From 19d6240dcc5e5c8bd6e1e3c588b92e837af76f9e Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Sat, 10 Dec 2022 16:28:05 +0000 Subject: [PATCH 0001/1291] feat(algebraic_topology/dold-kan): The Dold-Kan equivalence for additive categories (#17640) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR defines the Dold-Kan equivalence `karoubi (simplicial_object C) ≌ karoubi (chain_complex C ℕ)` between the idempotent completions of the categories `simplicial_object C` and `chain_complex C ℕ` when `C` is an additive category. --- .../dold_kan/equivalence_additive.lean | 62 +++++++++++++ .../dold_kan/n_comp_gamma.lean | 90 ++++++++++++++++++- 2 files changed, 151 insertions(+), 1 deletion(-) create mode 100644 src/algebraic_topology/dold_kan/equivalence_additive.lean diff --git a/src/algebraic_topology/dold_kan/equivalence_additive.lean b/src/algebraic_topology/dold_kan/equivalence_additive.lean new file mode 100644 index 0000000000000..166b18d6bce1d --- /dev/null +++ b/src/algebraic_topology/dold_kan/equivalence_additive.lean @@ -0,0 +1,62 @@ +/- +Copyright (c) 2022 Joël Riou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joël Riou +-/ + +import algebraic_topology.dold_kan.n_comp_gamma + +/-! The Dold-Kan equivalence for additive categories. + +This file defines `preadditive.dold_kan.equivalence` which is the equivalence +of categories `karoubi (simplicial_object C) ≌ karoubi (chain_complex C ℕ)`. + +-/ + +noncomputable theory + +open category_theory category_theory.category category_theory.limits + category_theory.idempotents algebraic_topology.dold_kan + +variables {C : Type*} [category C] [preadditive C] + +namespace category_theory + +namespace preadditive + +namespace dold_kan + +/-- The functor `karoubi (simplicial_object C) ⥤ karoubi (chain_complex C ℕ)` of +the Dold-Kan equivalence for additive categories. -/ +@[simps] +def N : karoubi (simplicial_object C) ⥤ karoubi (chain_complex C ℕ) := N₂ + +variable [has_finite_coproducts C] + +/-- The inverse functor `karoubi (chain_complex C ℕ) ⥤ karoubi (simplicial_object C)` of +the Dold-Kan equivalence for additive categories. -/ +@[simps] +def Γ : karoubi (chain_complex C ℕ) ⥤ karoubi (simplicial_object C) := Γ₂ + +/-- The Dold-Kan equivalence `karoubi (simplicial_object C) ≌ karoubi (chain_complex C ℕ)` +for additive categories. -/ +@[simps] +def equivalence : karoubi (simplicial_object C) ≌ karoubi (chain_complex C ℕ) := +{ functor := N, + inverse := Γ, + unit_iso := Γ₂N₂, + counit_iso := N₂Γ₂, + functor_unit_iso_comp' := λ P, begin + let α := N.map_iso (Γ₂N₂.app P), + let β := N₂Γ₂.app (N.obj P), + symmetry, + change 𝟙 _ = α.hom ≫ β.hom, + rw [← iso.inv_comp_eq, comp_id, ← comp_id β.hom, ← iso.inv_comp_eq], + exact algebraic_topology.dold_kan.identity_N₂_objectwise P, + end } + +end dold_kan + +end preadditive + +end category_theory diff --git a/src/algebraic_topology/dold_kan/n_comp_gamma.lean b/src/algebraic_topology/dold_kan/n_comp_gamma.lean index 816749fc7e191..415233b39b072 100644 --- a/src/algebraic_topology/dold_kan/n_comp_gamma.lean +++ b/src/algebraic_topology/dold_kan/n_comp_gamma.lean @@ -12,7 +12,7 @@ import algebraic_topology.dold_kan.n_reflects_iso In order to construct the unit isomorphism of the Dold-Kan equivalence, we first construct natural transformations `Γ₂N₁.nat_trans : N₁ ⋙ Γ₂ ⟶ to_karoubi (simplicial_object C)` and -`Γ₂N₂.nat_trans : N₂ ⋙ Γ₂ ⟶ 𝟭 (simplicial_object C)` (TODO). +`Γ₂N₂.nat_trans : N₂ ⋙ Γ₂ ⟶ 𝟭 (simplicial_object C)`. It is then shown that `Γ₂N₂.nat_trans` is an isomorphism by using that it becomes an isomorphism after the application of the functor `N₂ : karoubi (simplicial_object C) ⥤ karoubi (chain_complex C ℕ)` @@ -158,6 +158,94 @@ def nat_trans : (N₁ : simplicial_object C ⥤ _) ⋙ Γ₂ ⟶ to_karoubi _ := end Γ₂N₁ +/-- The compatibility isomorphism relating `N₂ ⋙ Γ₂` and `N₁ ⋙ Γ₂`. -/ +@[simps] +def compatibility_Γ₂N₁_Γ₂N₂ : to_karoubi (simplicial_object C) ⋙ N₂ ⋙ Γ₂ ≅ N₁ ⋙ Γ₂ := +eq_to_iso (functor.congr_obj (functor_extension₁_comp_whiskering_left_to_karoubi _ _) (N₁ ⋙ Γ₂)) + +namespace Γ₂N₂ + +/-- The natural transformation `N₂ ⋙ Γ₂ ⟶ 𝟭 (simplicial_object C)`. -/ +def nat_trans : (N₂ : karoubi (simplicial_object C) ⥤ _) ⋙ Γ₂ ⟶ 𝟭 _ := +((whiskering_left _ _ _).obj _).preimage (compatibility_Γ₂N₁_Γ₂N₂.hom ≫ Γ₂N₁.nat_trans) + +lemma nat_trans_app_f_app (P : karoubi (simplicial_object C)) : + Γ₂N₂.nat_trans.app P = (N₂ ⋙ Γ₂).map P.decomp_id_i ≫ + (compatibility_Γ₂N₁_Γ₂N₂.hom ≫ Γ₂N₁.nat_trans).app P.X ≫ P.decomp_id_p := +whiskering_left_obj_preimage_app ((compatibility_Γ₂N₁_Γ₂N₂.hom ≫ Γ₂N₁.nat_trans)) P + +end Γ₂N₂ + +lemma compatibility_Γ₂N₁_Γ₂N₂_nat_trans (X : simplicial_object C) : + Γ₂N₁.nat_trans.app X = (compatibility_Γ₂N₁_Γ₂N₂.app X).inv ≫ + Γ₂N₂.nat_trans.app ((to_karoubi _).obj X) := +begin + rw [← cancel_epi (compatibility_Γ₂N₁_Γ₂N₂.app X).hom, iso.hom_inv_id_assoc], + exact congr_app (((whiskering_left _ _ _).obj _).image_preimage + (compatibility_Γ₂N₁_Γ₂N₂.hom ≫ Γ₂N₁.nat_trans : _ ⟶ to_karoubi _ ⋙ 𝟭 _ )).symm X, +end + +lemma identity_N₂_objectwise (P : karoubi (simplicial_object C)) : + N₂Γ₂.inv.app (N₂.obj P) ≫ N₂.map (Γ₂N₂.nat_trans.app P) = 𝟙 (N₂.obj P) := +begin + ext n, + have eq₁ : (N₂Γ₂.inv.app (N₂.obj P)).f.f n = P_infty.f n ≫ P.p.app (op [n]) ≫ + (Γ₀.splitting (N₂.obj P).X).ι_summand (splitting.index_set.id (op [n])), + { simp only [N₂Γ₂_inv_app_f_f, N₂_obj_p_f, assoc], }, + have eq₂ : (Γ₀.splitting (N₂.obj P).X).ι_summand (splitting.index_set.id (op [n])) ≫ + (N₂.map (Γ₂N₂.nat_trans.app P)).f.f n = P_infty.f n ≫ P.p.app (op [n]), + { dsimp [N₂], + simp only [Γ₂N₂.nat_trans_app_f_app, P_infty_on_Γ₀_splitting_summand_eq_self_assoc, + functor.comp_map, compatibility_Γ₂N₁_Γ₂N₂_hom, nat_trans.comp_app, + eq_to_hom_app, assoc, karoubi.comp_f, karoubi.eq_to_hom_f, eq_to_hom_refl, comp_id, + karoubi.decomp_id_p_f, karoubi.comp_p_assoc, Γ₂_map_f_app, + N₂_map_f_f, karoubi.decomp_id_i_f, Γ₂N₁.nat_trans_app_f_app], + erw [splitting.ι_desc_assoc, assoc, assoc, splitting.ι_desc_assoc], + dsimp [splitting.index_set.id, splitting.index_set.e], + simp only [assoc, nat_trans.naturality, P_infty_f_naturality_assoc, + app_idem_assoc, P_infty_f_idem_assoc], + erw [P.X.map_id, comp_id], }, + simp only [karoubi.comp_f, homological_complex.comp_f, karoubi.id_eq, N₂_obj_p_f, assoc, + eq₁, eq₂, P_infty_f_naturality_assoc, app_idem, P_infty_f_idem_assoc], +end + +lemma identity_N₂ : + ((𝟙 (N₂ : karoubi (simplicial_object C) ⥤ _ ) ◫ N₂Γ₂.inv) ≫ + (Γ₂N₂.nat_trans ◫ 𝟙 N₂) : N₂ ⟶ N₂) = 𝟙 N₂ := +by { ext P : 2, dsimp, rw [Γ₂.map_id, N₂.map_id, comp_id, id_comp, identity_N₂_objectwise P], } + +instance : is_iso (Γ₂N₂.nat_trans : (N₂ : karoubi (simplicial_object C) ⥤ _ ) ⋙ _ ⟶ _) := +begin + haveI : ∀ (P : karoubi (simplicial_object C)), is_iso (Γ₂N₂.nat_trans.app P), + { intro P, + haveI : is_iso (N₂.map (Γ₂N₂.nat_trans.app P)), + { have h := identity_N₂_objectwise P, + erw hom_comp_eq_id at h, + rw h, + apply_instance, }, + exact is_iso_of_reflects_iso _ N₂, }, + apply nat_iso.is_iso_of_is_iso_app, +end + +instance : is_iso (Γ₂N₁.nat_trans : (N₁ : simplicial_object C ⥤ _ ) ⋙ _ ⟶ _) := +begin + haveI : ∀ (X : simplicial_object C), is_iso (Γ₂N₁.nat_trans.app X), + { intro X, + rw compatibility_Γ₂N₁_Γ₂N₂_nat_trans, + apply_instance, }, + apply nat_iso.is_iso_of_is_iso_app, +end + +/-- The unit isomorphism of the Dold-Kan equivalence. -/ +@[simp] +def Γ₂N₂ : 𝟭 _ ≅ (N₂ : karoubi (simplicial_object C) ⥤ _) ⋙ Γ₂ := +(as_iso Γ₂N₂.nat_trans).symm + +/-- The natural isomorphism `to_karoubi (simplicial_object C) ≅ N₁ ⋙ Γ₂`. -/ +@[simps] +def Γ₂N₁ : to_karoubi _ ≅ (N₁ : simplicial_object C ⥤ _) ⋙ Γ₂ := +(as_iso Γ₂N₁.nat_trans).symm + end dold_kan end algebraic_topology From eedb2810a94de94cd1f3b3e4f83fea228590be33 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Sat, 10 Dec 2022 16:28:06 +0000 Subject: [PATCH 0002/1291] chore(*): add mathlib4 synchronization comments (#17871) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following PRs: * `algebra.group.ulift`: https://github.com/leanprover-community/mathlib4/pull/906 * `algebra.group.with_one.basic`: https://github.com/leanprover-community/mathlib4/pull/922 * `algebra.group_with_zero.units.lemmas`: https://github.com/leanprover-community/mathlib4/pull/920 * `algebra.order.field.defs`: https://github.com/leanprover-community/mathlib4/pull/905 * `algebra.order.group.abs`: https://github.com/leanprover-community/mathlib4/pull/896 * `algebra.order.group.inj_surj`: https://github.com/leanprover-community/mathlib4/pull/916 * `algebra.order.group.type_tags`: https://github.com/leanprover-community/mathlib4/pull/921 * `algebra.order.monoid.with_top`: https://github.com/leanprover-community/mathlib4/pull/902 * `algebra.order.positive.ring`: https://github.com/leanprover-community/mathlib4/pull/911 * `algebra.order.ring.canonical`: https://github.com/leanprover-community/mathlib4/pull/905 * `algebra.order.ring.char_zero`: https://github.com/leanprover-community/mathlib4/pull/905 * `algebra.order.ring.defs`: https://github.com/leanprover-community/mathlib4/pull/905 * `algebra.order.ring.inj_surj`: https://github.com/leanprover-community/mathlib4/pull/917 * `algebra.ring.idempotents`: https://github.com/leanprover-community/mathlib4/pull/918 --- src/algebra/group/ulift.lean | 4 ++++ src/algebra/group/with_one/basic.lean | 4 ++++ src/algebra/group_with_zero/units/lemmas.lean | 4 ++++ src/algebra/order/field/defs.lean | 4 ++++ src/algebra/order/group/abs.lean | 4 ++++ src/algebra/order/group/inj_surj.lean | 4 ++++ src/algebra/order/group/type_tags.lean | 6 +++++- src/algebra/order/monoid/with_top.lean | 6 +++++- src/algebra/order/positive/ring.lean | 4 ++++ src/algebra/order/ring/canonical.lean | 4 ++++ src/algebra/order/ring/char_zero.lean | 4 ++++ src/algebra/order/ring/defs.lean | 4 ++++ src/algebra/order/ring/inj_surj.lean | 4 ++++ src/algebra/ring/idempotents.lean | 4 ++++ 14 files changed, 58 insertions(+), 2 deletions(-) diff --git a/src/algebra/group/ulift.lean b/src/algebra/group/ulift.lean index 2e65b7c46f8b0..bc65fe96a52a0 100644 --- a/src/algebra/group/ulift.lean +++ b/src/algebra/group/ulift.lean @@ -10,6 +10,10 @@ import algebra.group_with_zero.inj_surj /-! # `ulift` instances for groups and monoids +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/906 +> Any changes to this file require a corresponding PR to mathlib4. + This file defines instances for group, monoid, semigroup and related structures on `ulift` types. (Recall `ulift α` is just a "copy" of a type `α` in a higher universe.) diff --git a/src/algebra/group/with_one/basic.lean b/src/algebra/group/with_one/basic.lean index 08b23c3badf5a..7ab93d206bf43 100644 --- a/src/algebra/group/with_one/basic.lean +++ b/src/algebra/group/with_one/basic.lean @@ -9,6 +9,10 @@ import algebra.hom.equiv.basic /-! # More operations on `with_one` and `with_zero` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/922 +> Any changes to this file require a corresponding PR to mathlib4. + This file defines various bundled morphisms on `with_one` and `with_zero` that were not available in `algebra/group/with_one/defs`. diff --git a/src/algebra/group_with_zero/units/lemmas.lean b/src/algebra/group_with_zero/units/lemmas.lean index b299f7ce0cffa..14e3d9e1879b4 100644 --- a/src/algebra/group_with_zero/units/lemmas.lean +++ b/src/algebra/group_with_zero/units/lemmas.lean @@ -10,6 +10,10 @@ import group_theory.group_action.units /-! # Further lemmas about units in a `monoid_with_zero` or a `group_with_zero`. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/920 +> Any changes to this file require a corresponding PR to mathlib4. + -/ variables {α M₀ G₀ M₀' G₀' F F' : Type*} diff --git a/src/algebra/order/field/defs.lean b/src/algebra/order/field/defs.lean index 58294ae15fda0..1c5bcde47f671 100644 --- a/src/algebra/order/field/defs.lean +++ b/src/algebra/order/field/defs.lean @@ -9,6 +9,10 @@ import algebra.order.ring.defs /-! # Linear ordered (semi)fields +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/905 +> Any changes to this file require a corresponding PR to mathlib4. + A linear ordered (semi)field is a (semi)field equipped with a linear order such that * addition respects the order: `a ≤ b → c + a ≤ c + b`; * multiplication of positives is positive: `0 < a → 0 < b → 0 < a * b`; diff --git a/src/algebra/order/group/abs.lean b/src/algebra/order/group/abs.lean index a105b0dd3f452..989793decc56c 100644 --- a/src/algebra/order/group/abs.lean +++ b/src/algebra/order/group/abs.lean @@ -9,6 +9,10 @@ import order.min_max /-! # Absolute values in ordered groups. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/896 +> Any changes to this file require a corresponding PR to mathlib4. -/ variables {α : Type*} diff --git a/src/algebra/order/group/inj_surj.lean b/src/algebra/order/group/inj_surj.lean index 363d340182f8c..78bfe0d1a91ce 100644 --- a/src/algebra/order/group/inj_surj.lean +++ b/src/algebra/order/group/inj_surj.lean @@ -9,6 +9,10 @@ import algebra.order.group.instances /-! # Pull back ordered groups along injective maps. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/916 +> Any changes to this file require a corresponding PR to mathlib4. -/ variables {α β : Type*} diff --git a/src/algebra/order/group/type_tags.lean b/src/algebra/order/group/type_tags.lean index 684ca9f8b0ccf..d6f1f3776fe1f 100644 --- a/src/algebra/order/group/type_tags.lean +++ b/src/algebra/order/group/type_tags.lean @@ -6,7 +6,11 @@ Authors: Jeremy Avigad, Leonardo de Moura, Mario Carneiro, Johannes Hölzl import algebra.order.group.instances import algebra.order.monoid.type_tags -/-! # Ordered group structures on `multiplicative α` and `additive α`. -/ +/-! # Ordered group structures on `multiplicative α` and `additive α`. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/921 +> Any changes to this file require a corresponding PR to mathlib4.-/ variables {α : Type*} diff --git a/src/algebra/order/monoid/with_top.lean b/src/algebra/order/monoid/with_top.lean index 8665adece4cc6..884be82cf244a 100644 --- a/src/algebra/order/monoid/with_top.lean +++ b/src/algebra/order/monoid/with_top.lean @@ -8,7 +8,11 @@ import algebra.order.monoid.order_dual import algebra.order.monoid.with_zero.basic import data.nat.cast.defs -/-! # Adjoining top/bottom elements to ordered monoids. -/ +/-! # Adjoining top/bottom elements to ordered monoids. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/902 +> Any changes to this file require a corresponding PR to mathlib4.-/ universes u v variables {α : Type u} {β : Type v} diff --git a/src/algebra/order/positive/ring.lean b/src/algebra/order/positive/ring.lean index 9b76e2a533f77..757201061142d 100644 --- a/src/algebra/order/positive/ring.lean +++ b/src/algebra/order/positive/ring.lean @@ -9,6 +9,10 @@ import algebra.ring.inj_surj /-! # Algebraic structures on the set of positive numbers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/911 +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define various instances (`add_semigroup`, `ordered_comm_monoid` etc) on the type `{x : R // 0 < x}`. In each case we try to require the weakest possible typeclass assumptions on `R` but possibly, there is a room for improvements. diff --git a/src/algebra/order/ring/canonical.lean b/src/algebra/order/ring/canonical.lean index 93b3f65cb756d..ca2016709b7f9 100644 --- a/src/algebra/order/ring/canonical.lean +++ b/src/algebra/order/ring/canonical.lean @@ -10,6 +10,10 @@ import group_theory.group_action.defs /-! # Canoncially ordered rings and semirings. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/905 +> Any changes to this file require a corresponding PR to mathlib4. + * `canonically_ordered_comm_semiring` - `canonically_ordered_add_monoid` & multiplication & `*` respects `≤` & no zero divisors - `comm_semiring` & `a ≤ b ↔ ∃ c, b = a + c` & no zero divisors diff --git a/src/algebra/order/ring/char_zero.lean b/src/algebra/order/ring/char_zero.lean index c0e3cdd3c0939..b222522e8345f 100644 --- a/src/algebra/order/ring/char_zero.lean +++ b/src/algebra/order/ring/char_zero.lean @@ -8,6 +8,10 @@ import algebra.order.ring.defs /-! # Strict ordered semiring have characteristic zero + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/905 +> Any changes to this file require a corresponding PR to mathlib4. -/ variables {α : Type*} diff --git a/src/algebra/order/ring/defs.lean b/src/algebra/order/ring/defs.lean index 612305ddff8a8..0b534d98e0973 100644 --- a/src/algebra/order/ring/defs.lean +++ b/src/algebra/order/ring/defs.lean @@ -19,6 +19,10 @@ import algebra.group.units /-! # Ordered rings and semirings +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/905 +> Any changes to this file require a corresponding PR to mathlib4. + This file develops the basics of ordered (semi)rings. Each typeclass here comprises diff --git a/src/algebra/order/ring/inj_surj.lean b/src/algebra/order/ring/inj_surj.lean index 881bd62bd6718..87cacef222eaf 100644 --- a/src/algebra/order/ring/inj_surj.lean +++ b/src/algebra/order/ring/inj_surj.lean @@ -10,6 +10,10 @@ import algebra.ring.inj_surj /-! # Pulling back ordered rings along injective maps. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/917 +> Any changes to this file require a corresponding PR to mathlib4. + -/ open function diff --git a/src/algebra/ring/idempotents.lean b/src/algebra/ring/idempotents.lean index 912b8de75f6de..c84331792fea4 100644 --- a/src/algebra/ring/idempotents.lean +++ b/src/algebra/ring/idempotents.lean @@ -11,6 +11,10 @@ import tactic.nth_rewrite /-! # Idempotents +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/918 +> Any changes to this file require a corresponding PR to mathlib4. + This file defines idempotents for an arbitary multiplication and proves some basic results, including: From 60fa54e778c9e85d930efae172435f42fb0d71f7 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sat, 10 Dec 2022 16:28:07 +0000 Subject: [PATCH 0003/1291] chore(number_theory/padics/padic_val): golf (#17885) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Also add `padic_val_nat_dvd_iff_le` and assume `≠ 0` instead of `0 <` in `pow_succ_padic_val_nat_not_dvd`. --- src/number_theory/padics/padic_val.lean | 94 +++++-------------- .../polynomial/cyclotomic/eval.lean | 2 +- 2 files changed, 27 insertions(+), 69 deletions(-) diff --git a/src/number_theory/padics/padic_val.lean b/src/number_theory/padics/padic_val.lean index b854224de87c8..67493b8b5f2bf 100644 --- a/src/number_theory/padics/padic_val.lean +++ b/src/number_theory/padics/padic_val.lean @@ -180,14 +180,8 @@ by simp `padic_val_rat_def`. -/ lemma padic_val_nat_def [hp : fact p.prime] {n : ℕ} (hn : 0 < n) : padic_val_nat p n - = (multiplicity p n).get (multiplicity.finite_nat_iff.2 ⟨nat.prime.ne_one hp.1, hn⟩) := -begin - simp [padic_val_nat], - split_ifs, - { refl }, - { exfalso, - exact h ⟨hp.out.ne_one, hn⟩ } -end + = (multiplicity p n).get (multiplicity.finite_nat_iff.2 ⟨hp.out.ne_one, hn⟩) := +dif_pos ⟨hp.out.ne_one, hn⟩ lemma padic_val_nat_def' {n : ℕ} (hp : p ≠ 1) (hn : 0 < n) : ↑(padic_val_nat p n) = multiplicity p n := @@ -198,16 +192,8 @@ by simp [padic_val_nat_def (fact.out p.prime).pos] lemma one_le_padic_val_nat_of_dvd {n : ℕ} [hp : fact p.prime] (hn : 0 < n) (div : p ∣ n) : 1 ≤ padic_val_nat p n := -begin - rw padic_val_nat_def hn, - let one_le_mul : _ ≤ multiplicity p n := - multiplicity.le_multiplicity_of_pow_dvd (by { rw [pow_one], exact div }), - simp only [nat.cast_one] at one_le_mul, - rcases one_le_mul with ⟨_, q⟩, - dsimp at q, - solve_by_elim, - exact hp -end +by rwa [← part_enat.coe_le_coe, padic_val_nat_def' hp.out.ne_one hn, ← pow_dvd_iff_le_multiplicity, + pow_one] lemma dvd_iff_padic_val_nat_ne_zero {p n : ℕ} [fact p.prime] (hn0 : n ≠ 0) : (p ∣ n) ↔ padic_val_nat p n ≠ 0 := @@ -364,16 +350,9 @@ variables {p a b : ℕ} [hp : fact p.prime] include hp /-- A rewrite lemma for `padic_val_nat p (a * b)` with conditions `a ≠ 0`, `b ≠ 0`. -/ -protected lemma mul (ha : a ≠ 0) (hb : b ≠ 0) : +protected lemma mul : a ≠ 0 → b ≠ 0 → padic_val_nat p (a * b) = padic_val_nat p a + padic_val_nat p b := -begin - apply int.coe_nat_inj, - simp only [padic_val_rat_of_nat, nat.cast_mul], - rw padic_val_rat.mul, - norm_cast, - exact cast_ne_zero.mpr ha, - exact cast_ne_zero.mpr hb -end +by exact_mod_cast @padic_val_rat.mul p _ a b protected lemma div_of_dvd (h : b ∣ a) : padic_val_nat p (a / b) = padic_val_nat p a - padic_val_nat p b := @@ -404,11 +383,15 @@ by rwa [padic_val_nat.pow _ (fact.out p.prime).ne_zero, padic_val_nat_self, mul_ protected lemma div_pow (dvd : p ^ a ∣ b) : padic_val_nat p (b / p ^ a) = (padic_val_nat p b) - a := begin - convert padic_val_nat.div_of_dvd dvd, - rw padic_val_nat.prime_pow, + rw [padic_val_nat.div_of_dvd dvd, padic_val_nat.prime_pow], exact hp end +protected lemma div' {m : ℕ} (cpm : coprime p m) {b : ℕ} (dvd : m ∣ b) : + padic_val_nat p (b / m) = padic_val_nat p b := +by rw [padic_val_nat.div_of_dvd dvd, eq_zero_of_not_dvd (hp.out.coprime_iff_not_dvd.mp cpm), + nat.sub_zero]; assumption + end padic_val_nat section padic_val_nat @@ -429,31 +412,24 @@ begin rw [multiplicity.pow_dvd_iff_le_multiplicity, padic_val_nat_def']; assumption end -lemma pow_succ_padic_val_nat_not_dvd {n : ℕ} [hp : fact p.prime] (hn : 0 < n) : - ¬ p ^ (padic_val_nat p n + 1) ∣ n := -begin - rw multiplicity.pow_dvd_iff_le_multiplicity, - rw padic_val_nat_def hn, - { rw [nat.cast_add, part_enat.coe_get], - simp only [nat.cast_one, not_le], - exact part_enat.lt_add_one (ne_top_iff_finite.mpr - (finite_nat_iff.mpr ⟨(fact.elim hp).ne_one, hn⟩)), }, - { apply_instance } -end +lemma padic_val_nat_dvd_iff_le [hp : fact p.prime] {a n : ℕ} (ha : a ≠ 0) : + p ^ n ∣ a ↔ n ≤ padic_val_nat p a := +by rw [pow_dvd_iff_le_multiplicity, ← padic_val_nat_def' hp.out.ne_one ha.bot_lt, + part_enat.coe_le_coe] lemma padic_val_nat_dvd_iff (n : ℕ) [hp : fact p.prime] (a : ℕ) : p ^ n ∣ a ↔ a = 0 ∨ n ≤ padic_val_nat p a := begin - split, - { rw [pow_dvd_iff_le_multiplicity, padic_val_nat], - split_ifs, - { rw part_enat.coe_le_iff, - exact λ hn, or.inr (hn _) }, - { simp only [true_and, not_lt, ne.def, not_false_iff, le_zero_iff, hp.out.ne_one] at h, - exact λ hn, or.inl h } }, - { rintro (rfl|h), - { exact dvd_zero (p ^ n) }, - { exact dvd_trans (pow_dvd_pow p h) pow_padic_val_nat_dvd } } + rcases eq_or_ne a 0 with rfl | ha, + { exact iff_of_true (dvd_zero _) (or.inl rfl) }, + { simp only [ha, false_or, padic_val_nat_dvd_iff_le ha] } +end + +lemma pow_succ_padic_val_nat_not_dvd {n : ℕ} [hp : fact p.prime] (hn : n ≠ 0) : + ¬ p ^ (padic_val_nat p n + 1) ∣ n := +begin + rw [padic_val_nat_dvd_iff_le hn, not_le], + exacts [nat.lt_succ_self _, hp] end lemma padic_val_nat_primes {q : ℕ} [hp : fact p.prime] [hq : fact q.prime] (neq : p ≠ q) : @@ -461,24 +437,6 @@ lemma padic_val_nat_primes {q : ℕ} [hp : fact p.prime] [hq : fact q.prime] (ne @padic_val_nat.eq_zero_of_not_dvd p q $ (not_congr (iff.symm (prime_dvd_prime_iff_eq hp.1 hq.1))).mp neq -protected lemma padic_val_nat.div' [hp : fact p.prime] : - ∀ {m : ℕ} (cpm : coprime p m) {b : ℕ} (dvd : m ∣ b), padic_val_nat p (b / m) = padic_val_nat p b -| 0 := λ cpm b dvd, by { rw zero_dvd_iff at dvd, rw [dvd, nat.zero_div] } -| (n + 1) := - λ cpm b dvd, - begin - rcases dvd with ⟨c, rfl⟩, - rw [mul_div_right c (nat.succ_pos _)],by_cases hc : c = 0, - { rw [hc, mul_zero] }, - { rw padic_val_nat.mul, - { suffices : ¬ p ∣ (n + 1), - { rw [padic_val_nat.eq_zero_of_not_dvd this, zero_add] }, - contrapose! cpm, - exact hp.1.dvd_iff_not_coprime.mp cpm }, - { exact nat.succ_ne_zero _ }, - { exact hc } } - end - open_locale big_operators lemma range_pow_padic_val_nat_subset_divisors {n : ℕ} (hn : n ≠ 0) : diff --git a/src/ring_theory/polynomial/cyclotomic/eval.lean b/src/ring_theory/polynomial/cyclotomic/eval.lean index 2866bc9cd568a..1da812833eafd 100644 --- a/src/ring_theory/polynomial/cyclotomic/eval.lean +++ b/src/ring_theory/polynomial/cyclotomic/eval.lean @@ -173,7 +173,7 @@ begin rw [finset.prod_singleton, ht, mul_left_comm, mul_comm, ←mul_assoc, mul_assoc] at this, have : (p ^ (padic_val_nat p n) * p : ℤ) ∣ n := ⟨_, this⟩, simp only [←pow_succ', ←int.nat_abs_dvd_iff_dvd, int.nat_abs_of_nat, int.nat_abs_pow] at this, - exact pow_succ_padic_val_nat_not_dvd hn' this, + exact pow_succ_padic_val_nat_not_dvd hn'.ne' this, { rintro x - y - hxy, apply nat.succ_injective, exact nat.pow_right_injective hp.two_le hxy } From 68f41d3ce8925a1e63d7a516648a532113c98280 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sat, 10 Dec 2022 16:28:08 +0000 Subject: [PATCH 0004/1291] chore(number_theory/ramification_inertia): make an argument explicit (#17890) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This makes it slightly easier to use the `ideal.factors.pi_quotient_linear_equiv` definition, as otherwise all the typeclass search on properties of `S` is postponed till after the proof argument `hp : map (algebra_map R S) p ≠ ⊥` is provided. --- src/number_theory/ramification_inertia.lean | 6 +++++- 1 file changed, 5 insertions(+), 1 deletion(-) diff --git a/src/number_theory/ramification_inertia.lean b/src/number_theory/ramification_inertia.lean index 056498830a63f..cd02289a4ccc1 100644 --- a/src/number_theory/ramification_inertia.lean +++ b/src/number_theory/ramification_inertia.lean @@ -763,6 +763,8 @@ rfl λ P, ideal.quotient.mk _ (algebra_map _ _ x) := rfl +variables (S) + /-- **Chinese remainder theorem** for a ring of integers: if the prime ideal `p : ideal R` factors in `S` as `∏ i, P i ^ e i`, then `S ⧸ I` factors `R ⧸ I`-linearly as `Π i, R ⧸ (P i ^ e i)`. -/ @@ -781,6 +783,8 @@ noncomputable def factors.pi_quotient_linear_equiv end, .. factors.pi_quotient_equiv p hp } +variables {S} + open_locale big_operators /-- The **fundamental identity** of ramification index `e` and inertia degree `f`: @@ -817,7 +821,7 @@ begin { rw ← finset.sum_attach, refine finset.sum_congr rfl (λ P _, _), rw factors.finrank_pow_ramification_idx }, - { refine linear_equiv.finrank_eq (factors.pi_quotient_linear_equiv p _).symm, + { refine linear_equiv.finrank_eq (factors.pi_quotient_linear_equiv S p _).symm, rwa [ne.def, ideal.map_eq_bot_iff_le_ker, (ring_hom.injective_iff_ker_eq_bot _).mp inj_RS, le_bot_iff] }, { exact finrank_quotient_map p K L }, From 955890d374bb632cd662a1ca51a5a6a11c39578f Mon Sep 17 00:00:00 2001 From: negiizhao Date: Sat, 10 Dec 2022 16:28:12 +0000 Subject: [PATCH 0005/1291] chore(data/finset): remove unused imports (#17891) Co-authored-by: Eric Wieser --- src/data/finset/basic.lean | 2 -- src/data/finset/image.lean | 2 ++ 2 files changed, 2 insertions(+), 2 deletions(-) diff --git a/src/data/finset/basic.lean b/src/data/finset/basic.lean index 75bba419c8e3b..79af265442f67 100644 --- a/src/data/finset/basic.lean +++ b/src/data/finset/basic.lean @@ -3,9 +3,7 @@ Copyright (c) 2015 Microsoft Corporation. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura, Jeremy Avigad, Minchao Wu, Mario Carneiro -/ -import data.int.order.basic import data.multiset.finset_ops -import algebra.hom.embedding import tactic.apply import tactic.nth_rewrite import tactic.monotonicity diff --git a/src/data/finset/image.lean b/src/data/finset/image.lean index 35b72e8a79da7..aae0100e070f7 100644 --- a/src/data/finset/image.lean +++ b/src/data/finset/image.lean @@ -3,7 +3,9 @@ Copyright (c) 2015 Microsoft Corporation. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura, Jeremy Avigad, Minchao Wu, Mario Carneiro -/ +import algebra.hom.embedding import data.finset.basic +import data.int.order.basic /-! # Image and map operations on finite sets From 2e0975f6a25dd3fbfb9e41556a77f075f6269748 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sat, 10 Dec 2022 19:33:30 +0000 Subject: [PATCH 0006/1291] refactor(algebra/group/defs): use `is_left_cancel_mul` etc (#17884) The Lean 4 version is here: leanprover-community/mathlib4#945. --- src/algebra/group/defs.lean | 136 +++++++++------------ src/algebra/group/type_tags.lean | 30 +++-- src/data/nat/modeq.lean | 10 +- src/data/nat/order/lemmas.lean | 2 +- src/linear_algebra/finite_dimensional.lean | 2 +- src/set_theory/game/impartial.lean | 4 +- src/set_theory/ordinal/natural_ops.lean | 8 +- 7 files changed, 93 insertions(+), 99 deletions(-) diff --git a/src/algebra/group/defs.lean b/src/algebra/group/defs.lean index ae71d164e6ef6..564c7c0eef520 100644 --- a/src/algebra/group/defs.lean +++ b/src/algebra/group/defs.lean @@ -133,6 +133,57 @@ class is_cancel_add (G : Type u) [has_add G] attribute [to_additive] is_cancel_mul +section is_left_cancel_mul +variables [is_left_cancel_mul G] {a b c : G} + +@[to_additive] +lemma mul_left_cancel : a * b = a * c → b = c := +is_left_cancel_mul.mul_left_cancel a b c + +@[to_additive] +lemma mul_left_cancel_iff : a * b = a * c ↔ b = c := +⟨mul_left_cancel, congr_arg _⟩ + +@[to_additive] +theorem mul_right_injective (a : G) : function.injective ((*) a) := +λ b c, mul_left_cancel + +@[simp, to_additive] +theorem mul_right_inj (a : G) {b c : G} : a * b = a * c ↔ b = c := +(mul_right_injective a).eq_iff + +@[to_additive] +theorem mul_ne_mul_right (a : G) {b c : G} : a * b ≠ a * c ↔ b ≠ c := +(mul_right_injective a).ne_iff + +end is_left_cancel_mul + +section is_right_cancel_mul + +variables [is_right_cancel_mul G] {a b c : G} + +@[to_additive] +lemma mul_right_cancel : a * b = c * b → a = c := +is_right_cancel_mul.mul_right_cancel a b c + +@[to_additive] +lemma mul_right_cancel_iff : b * a = c * a ↔ b = c := +⟨mul_right_cancel, congr_arg _⟩ + +@[to_additive] +theorem mul_left_injective (a : G) : function.injective (λ x, x * a) := +λ b c, mul_right_cancel + +@[simp, to_additive] +theorem mul_left_inj (a : G) {b c : G} : b * a = c * a ↔ b = c := +(mul_left_injective a).eq_iff + +@[to_additive] +theorem mul_ne_mul_left (a : G) {b c : G} : b * a ≠ c * a ↔ b ≠ c := +(mul_left_injective a).ne_iff + +end is_right_cancel_mul + end has_mul /-- A semigroup is a type with an associative `(*)`. -/ @@ -185,11 +236,7 @@ instance comm_semigroup.to_is_commutative : is_commutative G (*) := `is_right_cancel_add G`."] lemma comm_semigroup.is_right_cancel_mul.to_is_left_cancel_mul (G : Type u) [comm_semigroup G] [is_right_cancel_mul G] : is_left_cancel_mul G := -{ mul_left_cancel := λ a b c h, - begin - rw [mul_comm a b, mul_comm a c] at h, - exact is_right_cancel_mul.mul_right_cancel _ _ _ h - end } +⟨λ a b c h, mul_right_cancel $ (mul_comm _ _).trans (h.trans $ mul_comm _ _)⟩ /-- Any `comm_semigroup G` that satisfies `is_left_cancel_mul G` also satisfies `is_right_cancel_mul G`. -/ @@ -198,11 +245,7 @@ lemma comm_semigroup.is_right_cancel_mul.to_is_left_cancel_mul (G : Type u) [com `is_left_cancel_add G`."] lemma comm_semigroup.is_left_cancel_mul.to_is_right_cancel_mul (G : Type u) [comm_semigroup G] [is_left_cancel_mul G] : is_right_cancel_mul G := -{ mul_right_cancel := λ a b c h, - begin - rw [mul_comm a b, mul_comm c b] at h, - exact is_left_cancel_mul.mul_left_cancel _ _ _ h - end } +⟨λ a b c h, mul_left_cancel $ (mul_comm _ _).trans (h.trans $ mul_comm _ _)⟩ /-- Any `comm_semigroup G` that satisfies `is_left_cancel_mul G` also satisfies `is_cancel_mul G`. -/ @@ -210,12 +253,7 @@ lemma comm_semigroup.is_left_cancel_mul.to_is_right_cancel_mul (G : Type u) [com that satisfies `is_left_cancel_add G` also satisfies `is_cancel_add G`."] lemma comm_semigroup.is_left_cancel_mul.to_is_cancel_mul (G : Type u) [comm_semigroup G] [is_left_cancel_mul G] : is_cancel_mul G := -{ mul_left_cancel := is_left_cancel_mul.mul_left_cancel, - mul_right_cancel := λ a b c h, - begin - rw [mul_comm a b, mul_comm c b] at h, - exact is_left_cancel_mul.mul_left_cancel _ _ _ h - end } +{ .. ‹is_left_cancel_mul G›, .. comm_semigroup.is_left_cancel_mul.to_is_right_cancel_mul G } /-- Any `comm_semigroup G` that satisfies `is_right_cancel_mul G` also satisfies `is_cancel_mul G`. -/ @@ -223,12 +261,7 @@ lemma comm_semigroup.is_left_cancel_mul.to_is_cancel_mul (G : Type u) [comm_semi that satisfies `is_right_cancel_add G` also satisfies `is_cancel_add G`."] lemma comm_semigroup.is_right_cancel_mul.to_is_cancel_mul (G : Type u) [comm_semigroup G] [is_right_cancel_mul G] : is_cancel_mul G := -{ mul_left_cancel := λ a b c h, - begin - rw [mul_comm a b, mul_comm a c] at h, - exact is_right_cancel_mul.mul_right_cancel _ _ _ h - end, - mul_right_cancel := is_right_cancel_mul.mul_right_cancel } +{ .. ‹is_right_cancel_mul G›, .. comm_semigroup.is_right_cancel_mul.to_is_left_cancel_mul G } end comm_semigroup @@ -243,37 +276,12 @@ class add_left_cancel_semigroup (G : Type u) extends add_semigroup G := (add_left_cancel : ∀ a b c : G, a + b = a + c → b = c) attribute [to_additive add_left_cancel_semigroup] left_cancel_semigroup -section left_cancel_semigroup -variables [left_cancel_semigroup G] {a b c : G} - -@[to_additive] -lemma mul_left_cancel : a * b = a * c → b = c := -left_cancel_semigroup.mul_left_cancel a b c - -@[to_additive] -lemma mul_left_cancel_iff : a * b = a * c ↔ b = c := -⟨mul_left_cancel, congr_arg _⟩ - -@[to_additive] -theorem mul_right_injective (a : G) : function.injective ((*) a) := -λ b c, mul_left_cancel - -@[simp, to_additive] -theorem mul_right_inj (a : G) {b c : G} : a * b = a * c ↔ b = c := -(mul_right_injective a).eq_iff - -@[to_additive] -theorem mul_ne_mul_right (a : G) {b c : G} : a * b ≠ a * c ↔ b ≠ c := -(mul_right_injective a).ne_iff - /-- Any `left_cancel_semigroup` satisfies `is_left_cancel_mul`. -/ @[priority 100, to_additive "Any `add_left_cancel_semigroup` satisfies `is_left_cancel_add`."] instance left_cancel_semigroup.to_is_left_cancel_mul (G : Type u) [left_cancel_semigroup G] : is_left_cancel_mul G := { mul_left_cancel := left_cancel_semigroup.mul_left_cancel } -end left_cancel_semigroup - /-- A `right_cancel_semigroup` is a semigroup such that `a * b = c * b` implies `a = c`. -/ @[protect_proj, ancestor semigroup, ext] class right_cancel_semigroup (G : Type u) extends semigroup G := @@ -286,37 +294,12 @@ class add_right_cancel_semigroup (G : Type u) extends add_semigroup G := (add_right_cancel : ∀ a b c : G, a + b = c + b → a = c) attribute [to_additive add_right_cancel_semigroup] right_cancel_semigroup -section right_cancel_semigroup -variables [right_cancel_semigroup G] {a b c : G} - -@[to_additive] -lemma mul_right_cancel : a * b = c * b → a = c := -right_cancel_semigroup.mul_right_cancel a b c - -@[to_additive] -lemma mul_right_cancel_iff : b * a = c * a ↔ b = c := -⟨mul_right_cancel, congr_arg _⟩ - -@[to_additive] -theorem mul_left_injective (a : G) : function.injective (λ x, x * a) := -λ b c, mul_right_cancel - -@[simp, to_additive] -theorem mul_left_inj (a : G) {b c : G} : b * a = c * a ↔ b = c := -(mul_left_injective a).eq_iff - -@[to_additive] -theorem mul_ne_mul_left (a : G) {b c : G} : b * a ≠ c * a ↔ b ≠ c := -(mul_left_injective a).ne_iff - /-- Any `right_cancel_semigroup` satisfies `is_right_cancel_mul`. -/ @[priority 100, to_additive "Any `add_right_cancel_semigroup` satisfies `is_right_cancel_add`."] -instance right_cancel_semigroup.to_is_right_cancel_mul (G : Type u) - [right_cancel_semigroup G] : is_right_cancel_mul G := +instance right_cancel_semigroup.to_is_right_cancel_mul (G : Type u) [right_cancel_semigroup G] : + is_right_cancel_mul G := { mul_right_cancel := right_cancel_semigroup.mul_right_cancel } -end right_cancel_semigroup - /-- Typeclass for expressing that a type `M` with multiplication and a one satisfies `1 * a = a` and `a * 1 = a` for all `a : M`. -/ @[ancestor has_one has_mul] @@ -362,8 +345,6 @@ instance mul_one_class.to_is_right_id : is_right_id M (*) 1 := end mul_one_class - - section variables {M : Type u} @@ -580,8 +561,7 @@ class cancel_comm_monoid (M : Type u) extends left_cancel_monoid M, comm_monoid @[priority 100, to_additive] -- see Note [lower instance priority] instance cancel_comm_monoid.to_cancel_monoid (M : Type u) [cancel_comm_monoid M] : cancel_monoid M := -{ mul_right_cancel := λ a b c h, mul_left_cancel $ by rw [mul_comm, h, mul_comm], - .. ‹cancel_comm_monoid M› } +{ .. ‹cancel_comm_monoid M›, .. comm_semigroup.is_left_cancel_mul.to_is_right_cancel_mul M } /-- Any `cancel_monoid M` satisfies `is_cancel_mul M`. -/ @[priority 100, to_additive "Any `add_cancel_monoid M` satisfies `is_cancel_add M`."] diff --git a/src/algebra/group/type_tags.lean b/src/algebra/group/type_tags.lean index b4b13f9c76de3..5943f2a83c07e 100644 --- a/src/algebra/group/type_tags.lean +++ b/src/algebra/group/type_tags.lean @@ -124,21 +124,35 @@ instance [add_comm_semigroup α] : comm_semigroup (multiplicative α) := { mul_comm := @add_comm _ _, ..multiplicative.semigroup } +instance [has_mul α] [is_left_cancel_mul α] : is_left_cancel_add (additive α) := +{ add_left_cancel := @mul_left_cancel α _ _ } + +instance [has_add α] [is_left_cancel_add α] : is_left_cancel_mul (multiplicative α) := +{ mul_left_cancel := @add_left_cancel α _ _ } + +instance [has_mul α] [is_right_cancel_mul α] : is_right_cancel_add (additive α) := +{ add_right_cancel := @mul_right_cancel α _ _ } + +instance [has_add α] [is_right_cancel_add α] : is_right_cancel_mul (multiplicative α) := +{ mul_right_cancel := @add_right_cancel α _ _ } + +instance [has_mul α] [is_cancel_mul α] : is_cancel_add (additive α) := +{ ..additive.is_left_cancel_add, ..additive.is_right_cancel_add } + +instance [has_add α] [is_cancel_add α] : is_cancel_mul (multiplicative α) := +{ ..multiplicative.is_left_cancel_mul, ..multiplicative.is_right_cancel_mul } + instance [left_cancel_semigroup α] : add_left_cancel_semigroup (additive α) := -{ add_left_cancel := @mul_left_cancel _ _, - ..additive.add_semigroup } +{ ..additive.add_semigroup, ..additive.is_left_cancel_add } instance [add_left_cancel_semigroup α] : left_cancel_semigroup (multiplicative α) := -{ mul_left_cancel := @add_left_cancel _ _, - ..multiplicative.semigroup } +{ ..multiplicative.semigroup, ..multiplicative.is_left_cancel_mul } instance [right_cancel_semigroup α] : add_right_cancel_semigroup (additive α) := -{ add_right_cancel := @mul_right_cancel _ _, - ..additive.add_semigroup } +{ ..additive.add_semigroup, ..additive.is_right_cancel_add } instance [add_right_cancel_semigroup α] : right_cancel_semigroup (multiplicative α) := -{ mul_right_cancel := @add_right_cancel _ _, - ..multiplicative.semigroup } +{ ..multiplicative.semigroup, ..multiplicative.is_right_cancel_mul } instance [has_one α] : has_zero (additive α) := ⟨additive.of_mul 1⟩ diff --git a/src/data/nat/modeq.lean b/src/data/nat/modeq.lean index a31f6c9aaf642..59979347e8aaa 100644 --- a/src/data/nat/modeq.lean +++ b/src/data/nat/modeq.lean @@ -323,9 +323,9 @@ end lemma div_mod_eq_mod_mul_div (a b c : ℕ) : a / b % c = a % (b * c) / b := if hb0 : b = 0 then by simp [hb0] -else by rw [← @add_right_cancel_iff _ _ (c * (a / b / c)), mod_add_div, nat.div_div_eq_div_mul, - ← mul_right_inj' hb0, ← @add_left_cancel_iff _ _ (a % b), mod_add_div, - mul_add, ← @add_left_cancel_iff _ _ (a % (b * c) % b), add_left_comm, +else by rw [← @add_right_cancel_iff _ _ _ (c * (a / b / c)), mod_add_div, nat.div_div_eq_div_mul, + ← mul_right_inj' hb0, ← @add_left_cancel_iff _ _ _ (a % b), mod_add_div, + mul_add, ← @add_left_cancel_iff _ _ _ (a % (b * c) % b), add_left_comm, ← add_assoc (a % (b * c) % b), mod_add_div, ← mul_assoc, mod_add_div, mod_mul_right_mod] lemma add_mod_add_ite (a b c : ℕ) : @@ -342,7 +342,7 @@ else exact add_lt_add (nat.mod_lt _ (nat.pos_of_ne_zero hc0)) (nat.mod_lt _ (nat.pos_of_ne_zero hc0))), have h0 : 0 < (a % c + b % c) / c, from nat.div_pos h (nat.pos_of_ne_zero hc0), - rw [← @add_right_cancel_iff _ _ (c * ((a % c + b % c) / c)), add_comm _ c, add_assoc, + rw [← @add_right_cancel_iff _ _ _ (c * ((a % c + b % c) / c)), add_comm _ c, add_assoc, mod_add_div, le_antisymm (le_of_lt_succ h2) h0, mul_one, add_comm] }, { rw [nat.mod_eq_of_lt (lt_of_not_ge h), add_zero] } end @@ -358,7 +358,7 @@ by rw [← add_mod_add_ite, if_pos hc] lemma add_div {a b c : ℕ} (hc0 : 0 < c) : (a + b) / c = a / c + b / c + if c ≤ a % c + b % c then 1 else 0 := begin - rw [← mul_right_inj' hc0.ne', ← @add_left_cancel_iff _ _ ((a + b) % c + a % c + b % c)], + rw [← mul_right_inj' hc0.ne', ← @add_left_cancel_iff _ _ _ ((a + b) % c + a % c + b % c)], suffices : (a + b) % c + c * ((a + b) / c) + a % c + b % c = a % c + c * (a / c) + (b % c + c * (b / c)) + c * (if c ≤ a % c + b % c then 1 else 0) + (a + b) % c, diff --git a/src/data/nat/order/lemmas.lean b/src/data/nat/order/lemmas.lean index 1cb7cc86db405..b9e86827bdd55 100644 --- a/src/data/nat/order/lemmas.lean +++ b/src/data/nat/order/lemmas.lean @@ -61,7 +61,7 @@ end protected lemma div_eq_zero_iff {a b : ℕ} (hb : 0 < b) : a / b = 0 ↔ a < b := ⟨λ h, by rw [← mod_add_div a b, h, mul_zero, add_zero]; exact mod_lt _ hb, - λ h, by rw [← mul_right_inj' hb.ne', ← @add_left_cancel_iff _ _ (a % b), mod_add_div, + λ h, by rw [← mul_right_inj' hb.ne', ← @add_left_cancel_iff _ _ _ (a % b), mod_add_div, mod_eq_of_lt h, mul_zero, add_zero]⟩ protected lemma div_eq_zero {a b : ℕ} (hb : a < b) : a / b = 0 := diff --git a/src/linear_algebra/finite_dimensional.lean b/src/linear_algebra/finite_dimensional.lean index ba8e2b2484ea0..a8dc37a15b4b3 100644 --- a/src/linear_algebra/finite_dimensional.lean +++ b/src/linear_algebra/finite_dimensional.lean @@ -850,7 +850,7 @@ noncomputable def linear_equiv.quot_equiv_of_equiv (f₁ : p ≃ₗ[K] q) (f₂ : V ≃ₗ[K] V₂) : (V ⧸ p) ≃ₗ[K] (V₂ ⧸ q) := linear_equiv.of_finrank_eq _ _ begin - rw [← @add_right_cancel_iff _ _ (finrank K p), submodule.finrank_quotient_add_finrank, + rw [← @add_right_cancel_iff _ _ _ (finrank K p), submodule.finrank_quotient_add_finrank, linear_equiv.finrank_eq f₁, submodule.finrank_quotient_add_finrank, linear_equiv.finrank_eq f₂], end diff --git a/src/set_theory/game/impartial.lean b/src/set_theory/game/impartial.lean index f310757f0eaf4..94076d236cf67 100644 --- a/src/set_theory/game/impartial.lean +++ b/src/set_theory/game/impartial.lean @@ -143,11 +143,11 @@ lemma add_self : G + G ≈ 0 := /-- This lemma doesn't require `H` to be impartial. -/ lemma equiv_iff_add_equiv_zero (H : pgame) : H ≈ G ↔ H + G ≈ 0 := -by { rw [equiv_iff_game_eq, equiv_iff_game_eq, ←@add_right_cancel_iff _ _ (-⟦G⟧)], simpa } +by { rw [equiv_iff_game_eq, equiv_iff_game_eq, ←@add_right_cancel_iff _ _ _ (-⟦G⟧)], simpa } /-- This lemma doesn't require `H` to be impartial. -/ lemma equiv_iff_add_equiv_zero' (H : pgame) : G ≈ H ↔ G + H ≈ 0 := -by { rw [equiv_iff_game_eq, equiv_iff_game_eq, ←@add_left_cancel_iff _ _ (-⟦G⟧), eq_comm], simpa } +by { rw [equiv_iff_game_eq, equiv_iff_game_eq, ←@add_left_cancel_iff _ _ _ (-⟦G⟧), eq_comm], simpa } lemma le_zero_iff {G : pgame} [G.impartial] : G ≤ 0 ↔ 0 ≤ G := by rw [←zero_le_neg_iff, le_congr_right (neg_equiv_self G)] diff --git a/src/set_theory/ordinal/natural_ops.lean b/src/set_theory/ordinal/natural_ops.lean index 31a7f0ae3ebeb..526bc5ea3d49c 100644 --- a/src/set_theory/ordinal/natural_ops.lean +++ b/src/set_theory/ordinal/natural_ops.lean @@ -312,12 +312,12 @@ theorem nadd_le_nadd_iff_right : ∀ a {b c}, b ♯ a ≤ c ♯ a ↔ b ≤ c := @_root_.add_le_add_iff_right nat_ordinal _ _ _ _ theorem nadd_left_cancel : ∀ {a b c}, a ♯ b = a ♯ c → b = c := -@_root_.add_left_cancel nat_ordinal _ +@_root_.add_left_cancel nat_ordinal _ _ theorem nadd_right_cancel : ∀ {a b c}, a ♯ b = c ♯ b → a = c := -@_root_.add_right_cancel nat_ordinal _ +@_root_.add_right_cancel nat_ordinal _ _ theorem nadd_left_cancel_iff : ∀ {a b c}, a ♯ b = a ♯ c ↔ b = c := -@add_left_cancel_iff nat_ordinal _ +@add_left_cancel_iff nat_ordinal _ _ theorem nadd_right_cancel_iff : ∀ {a b c}, b ♯ a = c ♯ a ↔ b = c := -@add_right_cancel_iff nat_ordinal _ +@add_right_cancel_iff nat_ordinal _ _ end ordinal From 036348aee28cb6cf44001c1fba4887abc91ea482 Mon Sep 17 00:00:00 2001 From: negiizhao Date: Sat, 10 Dec 2022 19:33:31 +0000 Subject: [PATCH 0007/1291] chore(data/multiset/nodup): remove dependency on `multiset.powerset` (#17888) This flips the import direction between `data.multiset.powerset` and `data.multiset.nodup`, such that the former now imports the latter, moving lemmas as necessary. This matches how `data.list.sublists` (transitively) imports `data.list.nodup`. More importantly, this means that `finset` (which needs `multiset.nodup`) can be defined without needing the material on `list.sublists` and `multiset.powerset`. --- src/algebra/big_operators/basic.lean | 1 + src/data/finset/basic.lean | 5 +++++ src/data/finset/powerset.lean | 1 + src/data/multiset/nodup.lean | 15 +-------------- src/data/multiset/powerset.lean | 18 +++++++++++++++--- 5 files changed, 23 insertions(+), 17 deletions(-) diff --git a/src/algebra/big_operators/basic.lean b/src/algebra/big_operators/basic.lean index 65edbaf781e7b..98797855b4e12 100644 --- a/src/algebra/big_operators/basic.lean +++ b/src/algebra/big_operators/basic.lean @@ -12,6 +12,7 @@ import algebra.ring.opposite import data.finset.sum import data.fintype.basic import data.finset.sigma +import data.multiset.powerset import data.set.pairwise /-! diff --git a/src/data/finset/basic.lean b/src/data/finset/basic.lean index 79af265442f67..c06c20d3f9f90 100644 --- a/src/data/finset/basic.lean +++ b/src/data/finset/basic.lean @@ -2490,3 +2490,8 @@ lemma disjoint_to_finset_iff_disjoint : _root_.disjoint l.to_finset l'.to_finset multiset.disjoint_to_finset end list + +-- Assert that we define `finset` without the material on `list.sublists`. +-- Note that we cannot use `list.sublists` itself as that is defined very early. +assert_not_exists list.sublists_len +assert_not_exists multiset.powerset diff --git a/src/data/finset/powerset.lean b/src/data/finset/powerset.lean index f200fb029a76b..0dd929a6a0e97 100644 --- a/src/data/finset/powerset.lean +++ b/src/data/finset/powerset.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ import data.finset.lattice +import data.multiset.powerset /-! # The powerset of a finset diff --git a/src/data/multiset/nodup.lean b/src/data/multiset/nodup.lean index 73f744131d197..1d1b639604903 100644 --- a/src/data/multiset/nodup.lean +++ b/src/data/multiset/nodup.lean @@ -3,8 +3,8 @@ Copyright (c) 2015 Microsoft Corporation. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ +import data.list.nodup import data.multiset.bind -import data.multiset.powerset import data.multiset.range /-! @@ -161,19 +161,6 @@ nodup_of_le $ inter_le_right _ _ λ ⟨h₁, h₂⟩, nodup_iff_count_le_one.2 $ λ a, by rw [count_union]; exact max_le (nodup_iff_count_le_one.1 h₁ a) (nodup_iff_count_le_one.1 h₂ a)⟩ -@[simp] theorem nodup_powerset {s : multiset α} : nodup (powerset s) ↔ nodup s := -⟨λ h, (nodup_of_le (map_single_le_powerset _) h).of_map _, - quotient.induction_on s $ λ l h, - by simp; refine (nodup_sublists'.2 h).map_on _ ; exact - λ x sx y sy e, - (h.sublist_ext (mem_sublists'.1 sx) (mem_sublists'.1 sy)).1 - (quotient.exact e)⟩ - -alias nodup_powerset ↔ nodup.of_powerset nodup.powerset - -protected lemma nodup.powerset_len {n : ℕ} (h : nodup s) : nodup (powerset_len n s) := -nodup_of_le (powerset_len_le_powerset _ _) (nodup_powerset.2 h) - @[simp] lemma nodup_bind {s : multiset α} {t : α → multiset β} : nodup (bind s t) ↔ ((∀a∈s, nodup (t a)) ∧ (s.pairwise (λa b, disjoint (t a) (t b)))) := have h₁ : ∀a, ∃l:list β, t a = l, from diff --git a/src/data/multiset/powerset.lean b/src/data/multiset/powerset.lean index f5ce620045db5..9cf07ef3c2430 100644 --- a/src/data/multiset/powerset.lean +++ b/src/data/multiset/powerset.lean @@ -4,9 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ import data.list.sublists -import data.multiset.basic -import data.multiset.range -import data.multiset.bind +import data.multiset.nodup /-! # The powerset of a multiset @@ -270,4 +268,18 @@ begin exact coe_eq_coe.mpr ((list.range_bind_sublists_len_perm S).map _), end +@[simp] theorem nodup_powerset {s : multiset α} : nodup (powerset s) ↔ nodup s := +⟨λ h, (nodup_of_le (map_single_le_powerset _) h).of_map _, + quotient.induction_on s $ λ l h, + by simp; refine (nodup_sublists'.2 h).map_on _ ; exact + λ x sx y sy e, + (h.sublist_ext (mem_sublists'.1 sx) (mem_sublists'.1 sy)).1 + (quotient.exact e)⟩ + +alias nodup_powerset ↔ nodup.of_powerset nodup.powerset + +protected lemma nodup.powerset_len {n : ℕ} {s : multiset α} (h : nodup s) : + nodup (powerset_len n s) := +nodup_of_le (powerset_len_le_powerset _ _) (nodup_powerset.2 h) + end multiset From 4e87c8477c6c38b753f050bc9664b94ee859896c Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 10 Dec 2022 22:50:33 +0000 Subject: [PATCH 0008/1291] feat(algebra/module/big_operators): Product of sums (#17818) A few missing lemmas. `finset.sum_smul_sum` matches `finset.sum_mul_sum`. Unrelatingly, fix the copyright after the split in #17764. I list as authors: * Chris for #327 * Yury for #1910 * myself for this PR --- src/algebra/module/big_operators.lean | 26 +++++++++++++++++++------- src/data/multiset/bind.lean | 7 +++---- 2 files changed, 22 insertions(+), 11 deletions(-) diff --git a/src/algebra/module/big_operators.lean b/src/algebra/module/big_operators.lean index 4a918ebcfe987..2945a10c102fa 100644 --- a/src/algebra/module/big_operators.lean +++ b/src/algebra/module/big_operators.lean @@ -1,11 +1,10 @@ /- -Copyright (c) 2015 Nathaniel Thomas. All rights reserved. +Copyright (c) 2018 Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Nathaniel Thomas, Jeremy Avigad, Johannes Hölzl, Mario Carneiro +Authors: Chris Hughes, Yury Kudryashov, Yaël Dillies -/ - import algebra.module.basic -import algebra.big_operators.basic +import group_theory.group_action.big_operators /-! # Finite sums over modules over a ring @@ -13,20 +12,33 @@ import algebra.big_operators.basic open_locale big_operators -universes u v -variables {α R k S M M₂ M₃ ι : Type*} +variables {α β R M ι : Type*} + section add_comm_monoid variables [semiring R] [add_comm_monoid M] [module R M] (r s : R) (x y : M) -variables {R M} + lemma list.sum_smul {l : list R} {x : M} : l.sum • x = (l.map (λ r, r • x)).sum := ((smul_add_hom R M).flip x).map_list_sum l lemma multiset.sum_smul {l : multiset R} {x : M} : l.sum • x = (l.map (λ r, r • x)).sum := ((smul_add_hom R M).flip x).map_multiset_sum l +lemma multiset.sum_smul_sum {s : multiset R} {t : multiset M} : + s.sum • t.sum = ((s ×ˢ t).map $ λ p : R × M, p.fst • p.snd).sum := +begin + induction s using multiset.induction with a s ih, + { simp }, + { simp [add_smul, ih, ←multiset.smul_sum] } +end + lemma finset.sum_smul {f : ι → R} {s : finset ι} {x : M} : (∑ i in s, f i) • x = (∑ i in s, (f i) • x) := ((smul_add_hom R M).flip x).map_sum f s + +lemma finset.sum_smul_sum {f : α → R} {g : β → M} {s : finset α} {t : finset β} : + (∑ i in s, f i) • (∑ i in t, g i) = ∑ p in s ×ˢ t, f p.fst • g p.snd := +by { rw [finset.sum_product, finset.sum_smul, finset.sum_congr rfl], intros, rw finset.smul_sum } + end add_comm_monoid lemma finset.cast_card [comm_semiring R] (s : finset α) : (s.card : R) = ∑ a in s, 1 := diff --git a/src/data/multiset/bind.lean b/src/data/multiset/bind.lean index 3006a479e92ac..031bc820e400e 100644 --- a/src/data/multiset/bind.lean +++ b/src/data/multiset/bind.lean @@ -162,10 +162,9 @@ infixr (name := multiset.product) ` ×ˢ `:82 := multiset.product by { rw [product, list.product, ←coe_bind], simp } @[simp] lemma zero_product : @product α β 0 t = 0 := rfl ---TODO: Add `product_zero` - -@[simp] lemma cons_product : (a ::ₘ s) ×ˢ t = map (prod.mk a) t + s ×ˢ t := -by simp [product] +@[simp] lemma cons_product : (a ::ₘ s) ×ˢ t = map (prod.mk a) t + s ×ˢ t := by simp [product] +@[simp] lemma product_zero : s ×ˢ (0 : multiset β) = 0 := by simp [product] +@[simp] lemma product_cons : s ×ˢ (b ::ₘ t) = s.map (λ a, (a, b)) + s ×ˢ t := by simp [product] @[simp] lemma product_singleton : ({a} : multiset α) ×ˢ ({b} : multiset β) = {(a, b)} := by simp only [product, bind_singleton, map_singleton] From 43afc5ad87891456c57b5a183e3e617d67c2b1db Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Sun, 11 Dec 2022 08:23:16 +0000 Subject: [PATCH 0009/1291] refactor(topology/algebra/ring): rename subring.subring_topological_closure to subring.le_topological_closure (#17649) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Change the name of the different versions of the result `s ≤ s.closure` for various substructures `s`, eg. subring, subalgebra, etc., from `substructure.substructure_topological_closure` to `substructure.le_topological_structure`. Indeed, it was noted in [this discussion](https://github.com/leanprover-community/mathlib/pull/17212#discussion_r1018230281) that the original names do not conform to the general naming policy in mathlib. --- src/analysis/inner_product_space/l2_space.lean | 2 +- src/analysis/inner_product_space/projection.lean | 4 ++-- src/analysis/normed_space/star/gelfand_duality.lean | 2 +- src/topology/algebra/algebra.lean | 6 +++--- src/topology/algebra/group/basic.lean | 2 +- src/topology/algebra/module/basic.lean | 8 ++++---- src/topology/algebra/module/linear_pmap.lean | 2 +- src/topology/algebra/monoid.lean | 2 +- src/topology/algebra/ring.lean | 4 ++-- .../continuous_function/stone_weierstrass.lean | 12 ++++++------ 10 files changed, 22 insertions(+), 22 deletions(-) diff --git a/src/analysis/inner_product_space/l2_space.lean b/src/analysis/inner_product_space/l2_space.lean index 4d5d34a61afb9..9937cb02aa50a 100644 --- a/src/analysis/inner_product_space/l2_space.lean +++ b/src/analysis/inner_product_space/l2_space.lean @@ -365,7 +365,7 @@ begin cases b; exact orthogonal.complete_space K <|> assumption }, refine is_hilbert_sum.mk_internal _ K.orthogonal_family_self _, - refine le_trans _ (submodule.submodule_topological_closure _), + refine le_trans _ (submodule.le_topological_closure _), rw [supr_bool_eq, cond, cond], refine codisjoint.top_le _, exact submodule.is_compl_orthogonal_of_complete_space.codisjoint diff --git a/src/analysis/inner_product_space/projection.lean b/src/analysis/inner_product_space/projection.lean index 65841d853e12d..482ec4b838a5b 100644 --- a/src/analysis/inner_product_space/projection.lean +++ b/src/analysis/inner_product_space/projection.lean @@ -730,7 +730,7 @@ lemma submodule.orthogonal_orthogonal_eq_closure [complete_space E] : Kᗮᗮ = K.topological_closure := begin refine le_antisymm _ _, - { convert submodule.orthogonal_orthogonal_monotone K.submodule_topological_closure, + { convert submodule.orthogonal_orthogonal_monotone K.le_topological_closure, haveI : complete_space K.topological_closure := K.is_closed_topological_closure.complete_space_coe, rw K.topological_closure.orthogonal_orthogonal }, @@ -807,7 +807,7 @@ begin let y := (orthogonal_projection (⨆ i, U i).topological_closure x : E), have proj_x : ∀ i, orthogonal_projection (U i) x = orthogonal_projection (U i) y := λ i, (orthogonal_projection_orthogonal_projection_of_le - ((le_supr U i).trans (supr U).submodule_topological_closure) _).symm, + ((le_supr U i).trans (supr U).le_topological_closure) _).symm, suffices : ∀ ε > 0, ∃ I, ∀ i ≥ I, ‖(orthogonal_projection (U i) y : E) - y‖ < ε, { simpa only [proj_x, normed_add_comm_group.tendsto_at_top] using this }, intros ε hε, diff --git a/src/analysis/normed_space/star/gelfand_duality.lean b/src/analysis/normed_space/star/gelfand_duality.lean index b7096eee0179c..e391587d86287 100644 --- a/src/analysis/normed_space/star/gelfand_duality.lean +++ b/src/analysis/normed_space/star/gelfand_duality.lean @@ -148,7 +148,7 @@ begin have h : (gelfand_transform ℂ A).range.topological_closure = (gelfand_transform ℂ A).range, from le_antisymm (subalgebra.topological_closure_minimal _ le_rfl (gelfand_transform_isometry A).closed_embedding.closed_range) - (subalgebra.subalgebra_topological_closure _), + (subalgebra.le_topological_closure _), refine h ▸ continuous_map.subalgebra_is_R_or_C_topological_closure_eq_top_of_separates_points _ (λ _ _, _) (λ f hf, _), /- Separating points just means that elements of the `character_space` which agree at all points diff --git a/src/topology/algebra/algebra.lean b/src/topology/algebra/algebra.lean index 7bff5853a3d23..7057336c41ad8 100644 --- a/src/topology/algebra/algebra.lean +++ b/src/topology/algebra/algebra.lean @@ -80,7 +80,7 @@ variables [topological_semiring A] /-- The closure of a subalgebra in a topological algebra as a subalgebra. -/ def subalgebra.topological_closure (s : subalgebra R A) : subalgebra R A := { carrier := closure (s : set A), - algebra_map_mem' := λ r, s.to_subsemiring.subring_topological_closure (s.algebra_map_mem r), + algebra_map_mem' := λ r, s.to_subsemiring.le_topological_closure (s.algebra_map_mem r), .. s.to_subsemiring.topological_closure } @[simp] lemma subalgebra.topological_closure_coe (s : subalgebra R A) : @@ -90,7 +90,7 @@ rfl instance subalgebra.topological_semiring (s : subalgebra R A) : topological_semiring s := s.to_subsemiring.topological_semiring -lemma subalgebra.subalgebra_topological_closure (s : subalgebra R A) : +lemma subalgebra.le_topological_closure (s : subalgebra R A) : s ≤ s.topological_closure := subset_closure @@ -149,7 +149,7 @@ def algebra.elemental_algebra (x : A) : subalgebra R A := (algebra.adjoin R ({x} : set A)).topological_closure lemma algebra.self_mem_elemental_algebra (x : A) : x ∈ algebra.elemental_algebra R x := -set_like.le_def.mp (subalgebra.subalgebra_topological_closure (algebra.adjoin R ({x} : set A))) $ +set_like.le_def.mp (subalgebra.le_topological_closure (algebra.adjoin R ({x} : set A))) $ algebra.self_mem_adjoin_singleton R x variables {R} diff --git a/src/topology/algebra/group/basic.lean b/src/topology/algebra/group/basic.lean index 9abda0620f2db..9a7b147a6b09b 100644 --- a/src/topology/algebra/group/basic.lean +++ b/src/topology/algebra/group/basic.lean @@ -526,7 +526,7 @@ def subgroup.topological_closure (s : subgroup G) : subgroup G := (s.topological_closure : set G) = closure s := rfl -@[to_additive] lemma subgroup.subgroup_topological_closure (s : subgroup G) : +@[to_additive] lemma subgroup.le_topological_closure (s : subgroup G) : s ≤ s.topological_closure := subset_closure diff --git a/src/topology/algebra/module/basic.lean b/src/topology/algebra/module/basic.lean index c05e27069fbd8..014bc026c3e3b 100644 --- a/src/topology/algebra/module/basic.lean +++ b/src/topology/algebra/module/basic.lean @@ -209,7 +209,7 @@ def submodule.topological_closure (s : submodule R M) : submodule R M := (s.topological_closure : set M) = closure (s : set M) := rfl -lemma submodule.submodule_topological_closure (s : submodule R M) : +lemma submodule.le_topological_closure (s : submodule R M) : s ≤ s.topological_closure := subset_closure @@ -224,13 +224,13 @@ closure_minimal h ht lemma submodule.topological_closure_mono {s : submodule R M} {t : submodule R M} (h : s ≤ t) : s.topological_closure ≤ t.topological_closure := -s.topological_closure_minimal (h.trans t.submodule_topological_closure) +s.topological_closure_minimal (h.trans t.le_topological_closure) t.is_closed_topological_closure /-- The topological closure of a closed submodule `s` is equal to `s`. -/ lemma is_closed.submodule_topological_closure_eq {s : submodule R M} (hs : is_closed (s : set M)) : s.topological_closure = s := -le_antisymm (s.topological_closure_minimal rfl.le hs) s.submodule_topological_closure +le_antisymm (s.topological_closure_minimal rfl.le hs) s.le_topological_closure /-- A subspace is dense iff its topological closure is the entire space. -/ lemma submodule.dense_iff_topological_closure_eq_top {s : submodule R M} : @@ -246,7 +246,7 @@ is_closed_closure.complete_space_coe is either closed or dense. -/ lemma submodule.is_closed_or_dense_of_is_coatom (s : submodule R M) (hs : is_coatom s) : is_closed (s : set M) ∨ dense (s : set M) := -(hs.le_iff.mp s.submodule_topological_closure).swap.imp (is_closed_of_closure_subset ∘ eq.le) +(hs.le_iff.mp s.le_topological_closure).swap.imp (is_closed_of_closure_subset ∘ eq.le) submodule.dense_iff_topological_closure_eq_top.mpr lemma linear_map.is_closed_or_dense_ker [has_continuous_add M'] [is_simple_module R' R'] diff --git a/src/topology/algebra/module/linear_pmap.lean b/src/topology/algebra/module/linear_pmap.lean index 26777fabe0e9f..2740693517f52 100644 --- a/src/topology/algebra/module/linear_pmap.lean +++ b/src/topology/algebra/module/linear_pmap.lean @@ -122,7 +122,7 @@ begin by_cases hf : f.is_closable, { refine le_of_le_graph _, rw ←hf.graph_closure_eq_closure_graph, - exact (graph f).submodule_topological_closure }, + exact (graph f).le_topological_closure }, rw closure_def' hf, end diff --git a/src/topology/algebra/monoid.lean b/src/topology/algebra/monoid.lean index 289823cba9eaf..53a4e631d7a27 100644 --- a/src/topology/algebra/monoid.lean +++ b/src/topology/algebra/monoid.lean @@ -317,7 +317,7 @@ def submonoid.topological_closure (s : submonoid M) : submonoid M := mul_mem' := λ a b ha hb, s.top_closure_mul_self_subset ⟨a, b, ha, hb, rfl⟩ } @[to_additive] -lemma submonoid.submonoid_topological_closure (s : submonoid M) : +lemma submonoid.le_topological_closure (s : submonoid M) : s ≤ s.topological_closure := subset_closure diff --git a/src/topology/algebra/ring.lean b/src/topology/algebra/ring.lean index 82adab65fcd80..9ee8148180bc4 100644 --- a/src/topology/algebra/ring.lean +++ b/src/topology/algebra/ring.lean @@ -111,7 +111,7 @@ def subsemiring.topological_closure (s : subsemiring α) : subsemiring α := (s.topological_closure : set α) = closure (s : set α) := rfl -lemma subsemiring.subring_topological_closure (s : subsemiring α) : +lemma subsemiring.le_topological_closure (s : subsemiring α) : s ≤ s.topological_closure := subset_closure @@ -267,7 +267,7 @@ def subring.topological_closure (S : subring α) : subring α := ..S.to_submonoid.topological_closure, ..S.to_add_subgroup.topological_closure } -lemma subring.subring_topological_closure (s : subring α) : +lemma subring.le_topological_closure (s : subring α) : s ≤ s.topological_closure := subset_closure lemma subring.is_closed_topological_closure (s : subring α) : diff --git a/src/topology/continuous_function/stone_weierstrass.lean b/src/topology/continuous_function/stone_weierstrass.lean index 1c07645da0167..581de93c3648e 100644 --- a/src/topology/continuous_function/stone_weierstrass.lean +++ b/src/topology/continuous_function/stone_weierstrass.lean @@ -124,8 +124,8 @@ begin rw inf_eq, refine A.topological_closure.smul_mem (A.topological_closure.sub_mem - (A.topological_closure.add_mem (A.subalgebra_topological_closure f.property) - (A.subalgebra_topological_closure g.property)) _) _, + (A.topological_closure.add_mem (A.le_topological_closure f.property) + (A.le_topological_closure g.property)) _) _, exact_mod_cast abs_mem_subalgebra_closure A _, end @@ -145,8 +145,8 @@ begin rw sup_eq, refine A.topological_closure.smul_mem (A.topological_closure.add_mem - (A.topological_closure.add_mem (A.subalgebra_topological_closure f.property) - (A.subalgebra_topological_closure g.property)) _) _, + (A.topological_closure.add_mem (A.le_topological_closure f.property) + (A.le_topological_closure g.property)) _) _, exact_mod_cast abs_mem_subalgebra_closure A _, end @@ -296,13 +296,13 @@ begin apply set_like.ext', let L := A.topological_closure, have n : set.nonempty (L : set C(X, ℝ)) := - ⟨(1 : C(X, ℝ)), A.subalgebra_topological_closure A.one_mem⟩, + ⟨(1 : C(X, ℝ)), A.le_topological_closure A.one_mem⟩, convert sublattice_closure_eq_top (L : set C(X, ℝ)) n (λ f fm g gm, inf_mem_closed_subalgebra L A.is_closed_topological_closure ⟨f, fm⟩ ⟨g, gm⟩) (λ f fm g gm, sup_mem_closed_subalgebra L A.is_closed_topological_closure ⟨f, fm⟩ ⟨g, gm⟩) (subalgebra.separates_points.strongly - (subalgebra.separates_points_monotone (A.subalgebra_topological_closure) w)), + (subalgebra.separates_points_monotone (A.le_topological_closure) w)), { simp, }, end From 9925a11537df32df1cdc61d9b09d0d415c92eb1d Mon Sep 17 00:00:00 2001 From: negiizhao Date: Sun, 11 Dec 2022 10:34:22 +0000 Subject: [PATCH 0010/1291] chore(data/list/range): split & reduce imports (#17887) This PR splits most of the lemmas about `list.fin_range` into a new file. `list.fin_range` is much less useful than `list.range`, but we need to import `data.list.of_fn` for `list.fin_range`, and then `data.list.of_fn` imports `data.fin.tuple.basic` and `data.fin.tuple.basic` imports many things. --- src/algebra/big_operators/fin.lean | 1 + src/algebra/graded_monoid.lean | 2 +- src/algebra/order/hom/ring.lean | 1 + src/data/fin/tuple/sort.lean | 3 +- src/data/fin/vec_notation.lean | 1 + src/data/finset/image.lean | 1 + src/data/list/fin_range.lean | 87 +++++++++++++++++++++++ src/data/list/indexes.lean | 1 + src/data/list/nat_antidiagonal.lean | 1 + src/data/list/perm.lean | 21 ------ src/data/list/range.lean | 57 +-------------- src/data/list/sort.lean | 1 + src/data/matrix/dmatrix.lean | 1 + src/linear_algebra/multilinear/basic.lean | 1 + 14 files changed, 101 insertions(+), 78 deletions(-) create mode 100644 src/data/list/fin_range.lean diff --git a/src/algebra/big_operators/fin.lean b/src/algebra/big_operators/fin.lean index 4090639cea233..20af1963a7386 100644 --- a/src/algebra/big_operators/fin.lean +++ b/src/algebra/big_operators/fin.lean @@ -5,6 +5,7 @@ Authors: Yury Kudryashov, Anne Baanen -/ import data.fintype.big_operators import data.fintype.fin +import data.list.fin_range import logic.equiv.fin /-! diff --git a/src/algebra/graded_monoid.lean b/src/algebra/graded_monoid.lean index 9aae3d0e47614..a0c0af93134a6 100644 --- a/src/algebra/graded_monoid.lean +++ b/src/algebra/graded_monoid.lean @@ -5,7 +5,7 @@ Authors: Eric Wieser -/ import algebra.group.inj_surj import data.list.big_operators.basic -import data.list.range +import data.list.fin_range import group_theory.group_action.defs import group_theory.submonoid.basic import data.set_like.basic diff --git a/src/algebra/order/hom/ring.lean b/src/algebra/order/hom/ring.lean index 6ded8e4b229e6..6cc404b68cf77 100644 --- a/src/algebra/order/hom/ring.lean +++ b/src/algebra/order/hom/ring.lean @@ -7,6 +7,7 @@ import algebra.order.archimedean import algebra.order.hom.monoid import algebra.order.ring.defs import algebra.ring.equiv +import tactic.by_contra import tactic.wlog /-! diff --git a/src/data/fin/tuple/sort.lean b/src/data/fin/tuple/sort.lean index e9bd615cca290..38430667b8f63 100644 --- a/src/data/fin/tuple/sort.lean +++ b/src/data/fin/tuple/sort.lean @@ -3,9 +3,8 @@ Copyright (c) 2021 Kyle Miller. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kyle Miller -/ - -import data.fin.basic import data.finset.sort +import data.list.fin_range import data.prod.lex import group_theory.perm.basic diff --git a/src/data/fin/vec_notation.lean b/src/data/fin/vec_notation.lean index ce9dc706abd23..41a6acc587c32 100644 --- a/src/data/fin/vec_notation.lean +++ b/src/data/fin/vec_notation.lean @@ -3,6 +3,7 @@ Copyright (c) 2020 Anne Baanen. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Anne Baanen -/ +import data.fin.tuple.basic import data.list.range import group_theory.group_action.pi import meta.univs diff --git a/src/data/finset/image.lean b/src/data/finset/image.lean index aae0100e070f7..3696de6c34e8c 100644 --- a/src/data/finset/image.lean +++ b/src/data/finset/image.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura, Jeremy Avigad, Minchao Wu, Mario Carneiro -/ import algebra.hom.embedding +import data.fin.basic import data.finset.basic import data.int.order.basic diff --git a/src/data/list/fin_range.lean b/src/data/list/fin_range.lean new file mode 100644 index 0000000000000..0ae1a608ec1a4 --- /dev/null +++ b/src/data/list/fin_range.lean @@ -0,0 +1,87 @@ +/- +Copyright (c) 2018 Mario Carneiro. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Mario Carneiro, Kenny Lau, Scott Morrison +-/ +import data.list.of_fn +import data.list.perm + +/-! +# Lists of elements of `fin n` + +This file develops some results on `fin_range n`. +-/ + +universe u + +namespace list +variables {α : Type u} + +@[simp] lemma map_coe_fin_range (n : ℕ) : (fin_range n).map coe = list.range n := +begin + simp_rw [fin_range, map_pmap, fin.coe_mk, pmap_eq_map], + exact list.map_id _ +end + +lemma fin_range_succ_eq_map (n : ℕ) : + fin_range n.succ = 0 :: (fin_range n).map fin.succ := +begin + apply map_injective_iff.mpr fin.coe_injective, + rw [map_cons, map_coe_fin_range, range_succ_eq_map, fin.coe_zero, ←map_coe_fin_range, map_map, + map_map, function.comp, function.comp], + congr' 2 with x, + exact (fin.coe_succ _).symm, +end + +@[simp] lemma map_nth_le (l : list α) : + (fin_range l.length).map (λ n, l.nth_le n n.2) = l := +ext_le (by rw [length_map, length_fin_range]) $ λ n _ h, +by { rw ← nth_le_map_rev, congr, { rw nth_le_fin_range, refl }, { rw length_fin_range, exact h } } + +theorem of_fn_eq_pmap {α n} {f : fin n → α} : + of_fn f = pmap (λ i hi, f ⟨i, hi⟩) (range n) (λ _, mem_range.1) := +by rw [pmap_eq_map_attach]; from ext_le (by simp) + (λ i hi1 hi2, by { simp at hi1, simp [nth_le_of_fn f ⟨i, hi1⟩, -subtype.val_eq_coe] }) + +theorem of_fn_id (n) : of_fn id = fin_range n := of_fn_eq_pmap + +theorem of_fn_eq_map {α n} {f : fin n → α} : + of_fn f = (fin_range n).map f := +by rw [← of_fn_id, map_of_fn, function.right_id] + +theorem nodup_of_fn_of_injective {α n} {f : fin n → α} (hf : function.injective f) : + nodup (of_fn f) := +by { rw of_fn_eq_pmap, exact (nodup_range n).pmap (λ _ _ _ _ H, fin.veq_of_eq $ hf H) } + +theorem nodup_of_fn {α n} {f : fin n → α} : + nodup (of_fn f) ↔ function.injective f := +begin + refine ⟨_, nodup_of_fn_of_injective⟩, + refine fin.cons_induction _ (λ n x₀ xs ih, _) f, + { intro h, + exact function.injective_of_subsingleton _ }, + { intro h, + rw fin.cons_injective_iff, + simp_rw [of_fn_succ, fin.cons_succ, nodup_cons, fin.cons_zero, mem_of_fn] at h, + exact h.imp_right ih } +end + +end list + +open list + +lemma equiv.perm.map_fin_range_perm {n : ℕ} (σ : equiv.perm (fin n)) : + map σ (fin_range n) ~ fin_range n := +begin + rw [perm_ext ((nodup_fin_range n).map σ.injective) $ nodup_fin_range n], + simpa only [mem_map, mem_fin_range, true_and, iff_true] using σ.surjective +end + +/-- The list obtained from a permutation of a tuple `f` is permutation equivalent to +the list obtained from `f`. -/ +lemma equiv.perm.of_fn_comp_perm {n : ℕ} {α : Type u} (σ : equiv.perm (fin n)) (f : fin n → α) : + of_fn (f ∘ σ) ~ of_fn f := +begin + rw [of_fn_eq_map, of_fn_eq_map, ←map_map], + exact σ.map_fin_range_perm.map f, +end diff --git a/src/data/list/indexes.lean b/src/data/list/indexes.lean index bc47dbbec6dcf..99fda0119a68e 100644 --- a/src/data/list/indexes.lean +++ b/src/data/list/indexes.lean @@ -3,6 +3,7 @@ Copyright (c) 2020 Jannis Limperg. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jannis Limperg -/ +import data.list.of_fn import data.list.range /-! diff --git a/src/data/list/nat_antidiagonal.lean b/src/data/list/nat_antidiagonal.lean index 0ec67b3cdc549..de578eb557c0d 100644 --- a/src/data/list/nat_antidiagonal.lean +++ b/src/data/list/nat_antidiagonal.lean @@ -3,6 +3,7 @@ Copyright (c) 2019 Johan Commelin. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin -/ +import data.list.nodup import data.list.range /-! diff --git a/src/data/list/perm.lean b/src/data/list/perm.lean index 1dfe6214e7caa..af5455996a7de 100644 --- a/src/data/list/perm.lean +++ b/src/data/list/perm.lean @@ -4,12 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura, Jeremy Avigad, Mario Carneiro -/ import data.list.dedup -import data.list.lattice import data.list.permutation -import data.list.zip import data.list.range import data.nat.factorial.basic -import logic.relation /-! # List Permutations @@ -1343,21 +1340,3 @@ end end permutations end list - -open list - -lemma equiv.perm.map_fin_range_perm {n : ℕ} (σ : equiv.perm (fin n)) : - map σ (fin_range n) ~ fin_range n := -begin - rw [perm_ext ((nodup_fin_range n).map σ.injective) $ nodup_fin_range n], - simpa only [mem_map, mem_fin_range, true_and, iff_true] using σ.surjective -end - -/-- The list obtained from a permutation of a tuple `f` is permutation equivalent to -the list obtained from `f`. -/ -lemma equiv.perm.of_fn_comp_perm {n : ℕ} {α : Type uu} (σ : equiv.perm (fin n)) (f : fin n → α) : - of_fn (f ∘ σ) ~ of_fn f := -begin - rw [of_fn_eq_map, of_fn_eq_map, ←map_map], - exact σ.map_fin_range_perm.map f, -end diff --git a/src/data/list/range.lean b/src/data/list/range.lean index 2f48e9c6f5e2f..04e0bcf17b944 100644 --- a/src/data/list/range.lean +++ b/src/data/list/range.lean @@ -4,8 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Kenny Lau, Scott Morrison -/ import data.list.chain -import data.list.nodup -import data.list.of_fn import data.list.zip /-! @@ -187,7 +185,7 @@ theorem pairwise_gt_iota (n : ℕ) : pairwise (>) (iota n) := by simp only [iota_eq_reverse_range', pairwise_reverse, pairwise_lt_range'] theorem nodup_iota (n : ℕ) : nodup (iota n) := -by simp only [iota_eq_reverse_range', nodup_reverse, nodup_range'] +(pairwise_gt_iota n).imp (λ a b, ne_of_gt) theorem mem_iota {m n : ℕ} : m ∈ iota n ↔ 1 ≤ m ∧ m ≤ n := by simp only [iota_eq_reverse_range', mem_reverse, mem_range', add_comm, lt_succ_iff] @@ -209,10 +207,10 @@ def fin_range (n : ℕ) : list (fin n) := @[simp] lemma fin_range_zero : fin_range 0 = [] := rfl @[simp] lemma mem_fin_range {n : ℕ} (a : fin n) : a ∈ fin_range n := -mem_pmap.2 ⟨a.1, mem_range.2 a.2, fin.eta _ _⟩ +mem_pmap.2 ⟨a.1, mem_range.2 a.2, by { cases a, refl, }⟩ lemma nodup_fin_range (n : ℕ) : (fin_range n).nodup := -(nodup_range _).pmap $ λ _ _ _ _, fin.veq_of_eq +pairwise.pmap (nodup_range n) _ $ λ _ _ _ _, @fin.ne_of_vne _ ⟨_, _⟩ ⟨_, _⟩ @[simp] lemma length_fin_range (n : ℕ) : (fin_range n).length = n := by rw [fin_range, length_pmap, length_range] @@ -220,22 +218,6 @@ by rw [fin_range, length_pmap, length_range] @[simp] lemma fin_range_eq_nil {n : ℕ} : fin_range n = [] ↔ n = 0 := by rw [← length_eq_zero, length_fin_range] -@[simp] lemma map_coe_fin_range (n : ℕ) : (fin_range n).map coe = list.range n := -begin - simp_rw [fin_range, map_pmap, fin.coe_mk, pmap_eq_map], - exact list.map_id _ -end - -lemma fin_range_succ_eq_map (n : ℕ) : - fin_range n.succ = 0 :: (fin_range n).map fin.succ := -begin - apply map_injective_iff.mpr fin.coe_injective, - rw [map_cons, map_coe_fin_range, range_succ_eq_map, fin.coe_zero, ←map_coe_fin_range, map_map, - map_map, function.comp, function.comp], - congr' 2 with x, - exact (fin.coe_succ _).symm, -end - @[to_additive] theorem prod_range_succ {α : Type u} [monoid α] (f : ℕ → α) (n : ℕ) : ((range n.succ).map f).prod = ((range n).map f).prod * f n := @@ -285,37 +267,4 @@ option.some.inj $ by rw [← nth_le_nth _, nth_range (by simpa using H)] (fin_range n).nth_le i h = ⟨i, length_fin_range n ▸ h⟩ := by simp only [fin_range, nth_le_range, nth_le_pmap] -@[simp] lemma map_nth_le (l : list α) : - (fin_range l.length).map (λ n, l.nth_le n n.2) = l := -ext_le (by rw [length_map, length_fin_range]) $ λ n _ h, -by { rw ← nth_le_map_rev, congr, { rw nth_le_fin_range, refl }, { rw length_fin_range, exact h } } - -theorem of_fn_eq_pmap {α n} {f : fin n → α} : - of_fn f = pmap (λ i hi, f ⟨i, hi⟩) (range n) (λ _, mem_range.1) := -by rw [pmap_eq_map_attach]; from ext_le (by simp) - (λ i hi1 hi2, by { simp at hi1, simp [nth_le_of_fn f ⟨i, hi1⟩, -subtype.val_eq_coe] }) - -theorem of_fn_id (n) : of_fn id = fin_range n := of_fn_eq_pmap - -theorem of_fn_eq_map {α n} {f : fin n → α} : - of_fn f = (fin_range n).map f := -by rw [← of_fn_id, map_of_fn, function.right_id] - -theorem nodup_of_fn_of_injective {α n} {f : fin n → α} (hf : function.injective f) : - nodup (of_fn f) := -by { rw of_fn_eq_pmap, exact (nodup_range n).pmap (λ _ _ _ _ H, fin.veq_of_eq $ hf H) } - -theorem nodup_of_fn {α n} {f : fin n → α} : - nodup (of_fn f) ↔ function.injective f := -begin - refine ⟨_, nodup_of_fn_of_injective⟩, - refine fin.cons_induction _ (λ n x₀ xs ih, _) f, - { intro h, - exact function.injective_of_subsingleton _ }, - { intro h, - rw fin.cons_injective_iff, - simp_rw [of_fn_succ, fin.cons_succ, nodup_cons, fin.cons_zero, mem_of_fn] at h, - exact h.imp_right ih } -end - end list diff --git a/src/data/list/sort.lean b/src/data/list/sort.lean index 92ba2fed19d55..86373474b8734 100644 --- a/src/data/list/sort.lean +++ b/src/data/list/sort.lean @@ -3,6 +3,7 @@ Copyright (c) 2016 Jeremy Avigad. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad -/ +import data.list.of_fn import data.list.perm /-! diff --git a/src/data/matrix/dmatrix.lean b/src/data/matrix/dmatrix.lean index dacbb08124504..2fe70c3422ed0 100644 --- a/src/data/matrix/dmatrix.lean +++ b/src/data/matrix/dmatrix.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ +import algebra.group.pi import data.fintype.basic /-! diff --git a/src/linear_algebra/multilinear/basic.lean b/src/linear_algebra/multilinear/basic.lean index e3d22caa099e2..b683759e9b64f 100644 --- a/src/linear_algebra/multilinear/basic.lean +++ b/src/linear_algebra/multilinear/basic.lean @@ -7,6 +7,7 @@ import linear_algebra.basic import algebra.algebra.basic import algebra.big_operators.order import algebra.big_operators.ring +import data.list.fin_range import data.fintype.big_operators import data.fintype.sort From 4f4e7fa6ea235496f1e49d5230e7ca8a65c4df48 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sun, 11 Dec 2022 13:46:37 +0000 Subject: [PATCH 0011/1291] feat(group_theory/perm/basic): `subtype_perm` is a monoid hom (#17816) A few more `subtype_perm` lemmas. Clean up the `subtype_perm` section. Rename `equiv.perm.default_perm` to `equiv.perm.default_eq`. --- src/group_theory/perm/basic.lean | 89 ++++++++++++++++++-------- src/group_theory/perm/cycle/basic.lean | 2 +- src/group_theory/perm/sign.lean | 4 +- 3 files changed, 65 insertions(+), 30 deletions(-) diff --git a/src/group_theory/perm/basic.lean b/src/group_theory/perm/basic.lean index e3ee863f0aba0..d3e93cc4fb647 100644 --- a/src/group_theory/perm/basic.lean +++ b/src/group_theory/perm/basic.lean @@ -30,6 +30,8 @@ instance perm_group : group (perm α) := mul_one := refl_trans, mul_left_inv := self_trans_symm } +@[simp] lemma default_eq : (default : perm α) = 1 := rfl + /-- The permutation of a type is equivalent to the units group of the endomorphisms monoid of this type. -/ @[simps] def equiv_units_End : perm α ≃* units (function.End α) := @@ -234,29 +236,71 @@ lemma extend_domain_hom_injective : function.injective (extend_domain_hom f) := end extend_domain +section subtype +variables {p : α → Prop} {f : perm α} + /-- If the permutation `f` fixes the subtype `{x // p x}`, then this returns the permutation on `{x // p x}` induced by `f`. -/ -def subtype_perm (f : perm α) {p : α → Prop} (h : ∀ x, p x ↔ p (f x)) : perm {x // p x} := +def subtype_perm (f : perm α) (h : ∀ x, p x ↔ p (f x)) : perm {x // p x} := ⟨λ x, ⟨f x, (h _).1 x.2⟩, λ x, ⟨f⁻¹ x, (h (f⁻¹ x)).2 $ by simpa using x.2⟩, λ _, by simp only [perm.inv_apply_self, subtype.coe_eta, subtype.coe_mk], λ _, by simp only [perm.apply_inv_self, subtype.coe_eta, subtype.coe_mk]⟩ -@[simp] lemma subtype_perm_apply (f : perm α) {p : α → Prop} (h : ∀ x, p x ↔ p (f x)) +@[simp] lemma subtype_perm_apply (f : perm α) (h : ∀ x, p x ↔ p (f x)) (x : {x // p x}) : subtype_perm f h x = ⟨f x, (h _).1 x.2⟩ := rfl -@[simp] lemma subtype_perm_one (p : α → Prop) (h : ∀ x, p x ↔ p ((1 : perm α) x)) : - @subtype_perm α 1 p h = 1 := +@[simp] lemma subtype_perm_one (p : α → Prop) (h := λ _, iff.rfl) : @subtype_perm α p 1 h = 1 := equiv.ext $ λ ⟨_, _⟩, rfl +@[simp] lemma subtype_perm_mul (f g : perm α) (hf hg) : + (f.subtype_perm hf * g.subtype_perm hg : perm {x // p x}) = + (f * g).subtype_perm (λ x, (hg _).trans $ hf _) := rfl + +private lemma inv_aux : (∀ x, p x ↔ p (f x)) ↔ ∀ x, p x ↔ p (f⁻¹ x) := +f⁻¹.surjective.forall.trans $ by simp_rw [f.apply_inv_self, iff.comm] + +/-- See `equiv.perm.inv_subtype_perm`-/ +lemma subtype_perm_inv (f : perm α) (hf) : + f⁻¹.subtype_perm hf = (f.subtype_perm $ inv_aux.2 hf : perm {x // p x})⁻¹ := rfl + +/-- See `equiv.perm.subtype_perm_inv`-/ +@[simp] lemma inv_subtype_perm (f : perm α) (hf) : + (f.subtype_perm hf : perm {x // p x})⁻¹ = f⁻¹.subtype_perm (inv_aux.1 hf) := rfl + +private lemma pow_aux (hf : ∀ x, p x ↔ p (f x)) : ∀ {n : ℕ} x, p x ↔ p ((f ^ n) x) +| 0 x := iff.rfl +| (n + 1) x := (pow_aux _).trans (hf _) + +@[simp] lemma subtype_perm_pow (f : perm α) (n : ℕ) (hf) : + (f.subtype_perm hf : perm {x // p x}) ^ n = (f ^ n).subtype_perm (pow_aux hf) := +begin + induction n with n ih, + { simp }, + { simp_rw [pow_succ', ih, subtype_perm_mul] } +end + +private lemma zpow_aux (hf : ∀ x, p x ↔ p (f x)) : ∀ {n : ℤ} x, p x ↔ p ((f ^ n) x) +| (int.of_nat n) := pow_aux hf +| (int.neg_succ_of_nat n) := by { rw zpow_neg_succ_of_nat, exact inv_aux.1 (pow_aux hf) } + +@[simp] lemma subtype_perm_zpow (f : perm α) (n : ℤ) (hf) : + (f.subtype_perm hf ^ n : perm {x // p x}) = (f ^ n).subtype_perm (zpow_aux hf) := +begin + induction n with n ih, + { exact subtype_perm_pow _ _ _ }, + { simp only [zpow_neg_succ_of_nat, subtype_perm_pow, subtype_perm_inv] } +end + +variables [decidable_pred p] {a : α} + /-- The inclusion map of permutations on a subtype of `α` into permutations of `α`, fixing the other points. -/ -def of_subtype {p : α → Prop} [decidable_pred p] : perm (subtype p) →* perm α := +def of_subtype : perm (subtype p) →* perm α := { to_fun := λ f, extend_domain f (equiv.refl (subtype p)), map_one' := equiv.perm.extend_domain_one _, map_mul' := λ f g, (equiv.perm.extend_domain_mul _ f g).symm, } -lemma of_subtype_subtype_perm {f : perm α} {p : α → Prop} [decidable_pred p] - (h₁ : ∀ x, p x ↔ p (f x)) (h₂ : ∀ x, f x ≠ x → p x) : +lemma of_subtype_subtype_perm {f : perm α} (h₁ : ∀ x, p x ↔ p (f x)) (h₂ : ∀ x, f x ≠ x → p x) : of_subtype (subtype_perm f h₁) = f := equiv.ext $ λ x, begin by_cases hx : p x, @@ -266,32 +310,24 @@ equiv.ext $ λ x, begin { exact hx, }, } end -lemma of_subtype_apply_of_mem {p : α → Prop} [decidable_pred p] - (f : perm (subtype p)) {x : α} (hx : p x) : - of_subtype f x = f ⟨x, hx⟩ := extend_domain_apply_subtype f _ hx +lemma of_subtype_apply_of_mem (f : perm (subtype p)) (ha : p a) : of_subtype f a = f ⟨a, ha⟩ := +extend_domain_apply_subtype _ _ _ -@[simp] lemma of_subtype_apply_coe {p : α → Prop} [decidable_pred p] - (f : perm (subtype p)) (x : subtype p) : - of_subtype f x = f x := +@[simp] lemma of_subtype_apply_coe (f : perm (subtype p)) (x : subtype p) : of_subtype f x = f x := subtype.cases_on x $ λ _, of_subtype_apply_of_mem f -lemma of_subtype_apply_of_not_mem {p : α → Prop} [decidable_pred p] - (f : perm (subtype p)) {x : α} (hx : ¬ p x) : - of_subtype f x = x := extend_domain_apply_not_subtype f (equiv.refl (subtype p)) hx +lemma of_subtype_apply_of_not_mem (f : perm (subtype p)) (ha : ¬ p a) : of_subtype f a = a := +extend_domain_apply_not_subtype _ _ ha -lemma mem_iff_of_subtype_apply_mem {p : α → Prop} [decidable_pred p] - (f : perm (subtype p)) (x : α) : +lemma mem_iff_of_subtype_apply_mem (f : perm (subtype p)) (x : α) : p x ↔ p ((of_subtype f : α → α) x) := if h : p x then by simpa only [h, true_iff, monoid_hom.coe_mk, of_subtype_apply_of_mem f h] using (f ⟨x, h⟩).2 else by simp [h, of_subtype_apply_of_not_mem f h] -@[simp] lemma subtype_perm_of_subtype {p : α → Prop} [decidable_pred p] (f : perm (subtype p)) : +@[simp] lemma subtype_perm_of_subtype (f : perm (subtype p)) : subtype_perm (of_subtype f) (mem_iff_of_subtype_apply_mem f) = f := -equiv.ext $ λ ⟨x, hx⟩, - subtype.coe_injective (of_subtype_apply_of_mem f hx) - -@[simp] lemma default_perm {n : Type*} : (default : perm n) = 1 := rfl +equiv.ext $ λ x, subtype.coe_injective (of_subtype_apply_coe f x) /-- Permutations on a subtype are equivalent to permutations on the original type that fix pointwise the rest. -/ @@ -305,16 +341,15 @@ the rest. -/ right_inv := λ f, subtype.ext (equiv.perm.of_subtype_subtype_perm _ $ λ a, not.decidable_imp_symm $ f.prop a) } -lemma subtype_equiv_subtype_perm_apply_of_mem {α : Type*} {p : α → Prop} - [decidable_pred p] (f : perm (subtype p)) {a : α} (h : p a) : +lemma subtype_equiv_subtype_perm_apply_of_mem (f : perm (subtype p)) (h : p a) : perm.subtype_equiv_subtype_perm p f a = f ⟨a, h⟩ := f.of_subtype_apply_of_mem h -lemma subtype_equiv_subtype_perm_apply_of_not_mem {α : Type*} {p : α → Prop} - [decidable_pred p] (f : perm (subtype p)) {a : α} (h : ¬ p a) : +lemma subtype_equiv_subtype_perm_apply_of_not_mem (f : perm (subtype p)) (h : ¬ p a) : perm.subtype_equiv_subtype_perm p f a = a := f.of_subtype_apply_of_not_mem h +end subtype end perm section swap diff --git a/src/group_theory/perm/cycle/basic.lean b/src/group_theory/perm/cycle/basic.lean index 98220da5fc040..8b6811ff8d37e 100644 --- a/src/group_theory/perm/cycle/basic.lean +++ b/src/group_theory/perm/cycle/basic.lean @@ -577,7 +577,7 @@ end /-- `f.cycle_of x` is the cycle of the permutation `f` to which `x` belongs. -/ def cycle_of [fintype α] (f : perm α) (x : α) : perm α := -of_subtype (@subtype_perm _ f (same_cycle f x) (λ _, same_cycle_apply.symm)) +of_subtype (subtype_perm f (λ _, same_cycle_apply.symm) : perm {y // same_cycle f x y}) lemma cycle_of_apply [fintype α] (f : perm α) (x y : α) : cycle_of f x y = if same_cycle f x y then f y else y := diff --git a/src/group_theory/perm/sign.lean b/src/group_theory/perm/sign.lean index e89ae7400e6c3..8ac87f929df47 100644 --- a/src/group_theory/perm/sign.lean +++ b/src/group_theory/perm/sign.lean @@ -578,9 +578,9 @@ lemma sign_bij [decidable_eq β] [fintype β] (hi : ∀ x₁ x₂ hx₁ hx₂, i x₁ hx₁ = i x₂ hx₂ → x₁ = x₂) (hg : ∀ y, g y ≠ y → ∃ x hx, i x hx = y) : sign f = sign g := -calc sign f = sign (@subtype_perm _ f (λ x, f x ≠ x) (by simp)) : +calc sign f = sign (subtype_perm f $ by simp : perm {x // f x ≠ x}) : (sign_subtype_perm _ _ (λ _, id)).symm -... = sign (@subtype_perm _ g (λ x, g x ≠ x) (by simp)) : +... = sign (subtype_perm g $ by simp : perm {x // g x ≠ x}) : sign_eq_sign_of_equiv _ _ (equiv.of_bijective (λ x : {x // f x ≠ x}, (⟨i x.1 x.2, have f (f x) ≠ f x, from mt (λ h, f.injective h) x.2, From cf9386b56953fb40904843af98b7a80757bbe7f9 Mon Sep 17 00:00:00 2001 From: Michael Stoll Date: Sun, 11 Dec 2022 18:31:17 +0000 Subject: [PATCH 0012/1291] feat(data/list/basic): add lemma map_eq_repeat_iff (#17832) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This small PR adds the following API lemma for lists: ```lean lemma map_eq_repeat_iff {α β} {l : list α} {f : α → β} {b : β} : l.map f = repeat b l.length ↔ (∀ x ∈ l, f x = b) ``` See this [Zulip thread](https://leanprover.zulipchat.com/#narrow/stream/217875-Is-there-code-for-X.3F/topic/More.20general.20version.20of.20.60list.2Emap_const.60.3F/near/314118320). (It also fixes a harmless typo in the proof of `last_map`.) --- src/data/list/basic.lean | 26 ++++++++++++++++++-------- 1 file changed, 18 insertions(+), 8 deletions(-) diff --git a/src/data/list/basic.lean b/src/data/list/basic.lean index bbbb9ac72c1bd..a1213cece0149 100644 --- a/src/data/list/basic.lean +++ b/src/data/list/basic.lean @@ -506,13 +506,6 @@ lemma subset_singleton_iff {a : α} : ∀ L : list α, L ⊆ [a] ↔ ∃ n, L = exact ⟨n.succ, by simp [mem_singleton.mp h.1]⟩ end -@[simp] theorem map_const (l : list α) (b : β) : map (function.const α b) l = repeat b l.length := -by induction l; [refl, simp only [*, map]]; split; refl - -theorem eq_of_mem_map_const {b₁ b₂ : β} {l : list α} (h : b₁ ∈ map (function.const α b₂) l) : - b₁ = b₂ := -by rw map_const at h; exact eq_of_mem_repeat h - @[simp] theorem map_repeat (f : α → β) (a : α) (n) : map f (repeat a n) = repeat (f a) n := by induction n; [refl, simp only [*, repeat, map]]; split; refl @@ -1791,13 +1784,30 @@ by { induction as, { refl }, { simp! [*, apply_ite (map f)] } } lemma last_map (f : α → β) {l : list α} (hl : l ≠ []) : (l.map f).last (mt eq_nil_of_map_eq_nil hl) = f (l.last hl) := begin - induction l with l_ih l_tl l_ih, + induction l with l_hd l_tl l_ih, { apply (hl rfl).elim }, { cases l_tl, { simp }, { simpa using l_ih } } end +lemma map_eq_repeat_iff {l : list α} {f : α → β} {b : β} : + l.map f = repeat b l.length ↔ (∀ x ∈ l, f x = b) := +begin + induction l with x l' ih, + { simp only [repeat, length, not_mem_nil, is_empty.forall_iff, implies_true_iff, + map_nil, eq_self_iff_true], }, + { simp only [map, length, mem_cons_iff, forall_eq_or_imp, repeat_succ, and.congr_right_iff], + exact λ _, ih, } +end + +@[simp] theorem map_const (l : list α) (b : β) : map (function.const α b) l = repeat b l.length := +map_eq_repeat_iff.mpr (λ x _, rfl) + +theorem eq_of_mem_map_const {b₁ b₂ : β} {l : list α} (h : b₁ ∈ map (function.const α b₂) l) : + b₁ = b₂ := +by rw map_const at h; exact eq_of_mem_repeat h + /-! ### map₂ -/ theorem nil_map₂ (f : α → β → γ) (l : list β) : map₂ f [] l = [] := From 1c521b4fb909320eca16b2bb6f8b5b0490b1cb5e Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Mon, 12 Dec 2022 06:59:12 +0000 Subject: [PATCH 0013/1291] feat(topology): `zero_at_infty_continuous_map`s are `uniform_continuous` (#17767) split from #17719 Co-authored-by: loefflerd Co-authored-by: David Loeffler --- .../continuous_function/zero_at_infty.lean | 11 +++++- src/topology/uniform_space/basic.lean | 5 +++ src/topology/uniform_space/compact.lean | 34 +++++++++++++++++++ 3 files changed, 49 insertions(+), 1 deletion(-) diff --git a/src/topology/continuous_function/zero_at_infty.lean b/src/topology/continuous_function/zero_at_infty.lean index 20f7f763be3cb..947c71539b09c 100644 --- a/src/topology/continuous_function/zero_at_infty.lean +++ b/src/topology/continuous_function/zero_at_infty.lean @@ -283,9 +283,18 @@ instance {R : Type*} [semiring R] [non_unital_non_assoc_semiring β] [topologica rw [←smul_eq_mul, ←smul_eq_mul, smul_comm], end } - end algebraic_structure +section uniform + +variables [uniform_space β] [uniform_space γ] [has_zero γ] + [zero_at_infty_continuous_map_class F β γ] + +lemma uniform_continuous (f : F) : uniform_continuous (f : β → γ) := +(map_continuous f).uniform_continuous_of_zero_at_infty (zero_at_infty f) + +end uniform + /-! ### Metric structure When `β` is a metric space, then every element of `C₀(α, β)` is bounded, and so there is a natural diff --git a/src/topology/uniform_space/basic.lean b/src/topology/uniform_space/basic.lean index 22adcfd5e0379..972b23fb70b02 100644 --- a/src/topology/uniform_space/basic.lean +++ b/src/topology/uniform_space/basic.lean @@ -683,6 +683,11 @@ lemma mem_nhds_right (y : α) {s : set (α×α)} (h : s ∈ 𝓤 α) : {x : α | (x, y) ∈ s} ∈ 𝓝 y := mem_nhds_left _ (symm_le_uniformity h) +lemma exists_mem_nhds_ball_subset_of_mem_nhds {a : α} {U : set α} (h : U ∈ 𝓝 a) : + ∃ (V ∈ 𝓝 a) (t ∈ 𝓤 α), ∀ a' ∈ V, uniform_space.ball a' t ⊆ U := +let ⟨t, ht, htU⟩ := comp_mem_uniformity_sets (mem_nhds_uniformity_iff_right.1 h) in +⟨_, mem_nhds_left a ht, t, ht, λ a₁ h₁ a₂ h₂, @htU (a, a₂) ⟨a₁, h₁, h₂⟩ rfl⟩ + lemma is_compact.nhds_set_basis_uniformity {p : ι → Prop} {s : ι → set (α × α)} (hU : (𝓤 α).has_basis p s) {K : set α} (hK : is_compact K) : (𝓝ˢ K).has_basis p (λ i, ⋃ x ∈ K, ball x (s i)) := diff --git a/src/topology/uniform_space/compact.lean b/src/topology/uniform_space/compact.lean index ccdd3cad6874f..68806ee7d7932 100644 --- a/src/topology/uniform_space/compact.lean +++ b/src/topology/uniform_space/compact.lean @@ -202,6 +202,40 @@ begin exact compact_space.uniform_continuous_of_continuous hf, end +/-- If `s` is compact and `f` is continuous at all points of `s`, then `f` is +"uniformly continuous at the set `s`", i.e. `f x` is close to `f y` whenever `x ∈ s` and `y` is +close to `x` (even if `y` is not itself in `s`, so this is a stronger assertion than +`uniform_continuous_on s`). -/ +lemma is_compact.uniform_continuous_at_of_continuous_at {r : set (β × β)} {s : set α} + (hs : is_compact s) (f : α → β) (hf : ∀ a ∈ s, continuous_at f a) (hr : r ∈ 𝓤 β) : + {x : α × α | x.1 ∈ s → (f x.1, f x.2) ∈ r} ∈ 𝓤 α := +begin + obtain ⟨t, ht, htsymm, htr⟩ := comp_symm_mem_uniformity_sets hr, + choose U hU T hT hb using λ a ha, exists_mem_nhds_ball_subset_of_mem_nhds + ((hf a ha).preimage_mem_nhds $ mem_nhds_left _ ht), + obtain ⟨fs, hsU⟩ := hs.elim_nhds_subcover' U hU, + apply mem_of_superset ((bInter_finset_mem fs).2 $ λ a _, hT a a.2), + rintro ⟨a₁, a₂⟩ h h₁, + obtain ⟨a, ha, haU⟩ := set.mem_Union₂.1 (hsU h₁), + apply htr, + refine ⟨f a, htsymm.mk_mem_comm.1 (hb _ _ _ haU _), hb _ _ _ haU _⟩, + exacts [mem_ball_self _ (hT a a.2), mem_Inter₂.1 h a ha], +end + +lemma continuous.uniform_continuous_of_zero_at_infty {f : α → β} [has_zero β] + (h_cont : continuous f) (h_zero : tendsto f (cocompact α) (𝓝 0)) : uniform_continuous f := +uniform_continuous_def.2 $ λ r hr, begin + obtain ⟨t, ht, htsymm, htr⟩ := comp_symm_mem_uniformity_sets hr, + obtain ⟨s, hs, hst⟩ := mem_cocompact.1 (h_zero $ mem_nhds_left 0 ht), + apply mem_of_superset (symmetrize_mem_uniformity $ hs.uniform_continuous_at_of_continuous_at + f (λ _ _, h_cont.continuous_at) $ symmetrize_mem_uniformity hr), + rintro ⟨b₁, b₂⟩ h, + by_cases h₁ : b₁ ∈ s, { exact (h.1 h₁).1 }, + by_cases h₂ : b₂ ∈ s, { exact (h.2 h₂).2 }, + apply htr, + exact ⟨0, htsymm.mk_mem_comm.1 (hst h₁), hst h₂⟩, +end + /-- A family of functions `α → β → γ` tends uniformly to its value at `x` if `α` is locally compact, `β` is compact and `f` is continuous on `U × (univ : set β)` for some neighborhood `U` of `x`. -/ lemma continuous_on.tendsto_uniformly [locally_compact_space α] [compact_space β] From 800f52292ec9385fdd2466f66d288d07c39cb7e3 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 12 Dec 2022 09:04:19 +0000 Subject: [PATCH 0014/1291] feat(analysis/convex/cone/basic): submodules are convex cones (#17749) --- src/analysis/convex/cone/basic.lean | 43 +++++++++++++++++++++++++++++ 1 file changed, 43 insertions(+) diff --git a/src/analysis/convex/cone/basic.lean b/src/analysis/convex/cone/basic.lean index 34135c4216f23..599fea0333854 100644 --- a/src/analysis/convex/cone/basic.lean +++ b/src/analysis/convex/cone/basic.lean @@ -394,6 +394,49 @@ instance : add_comm_semigroup (convex_cone 𝕜 E) := end module end ordered_semiring +end convex_cone + +namespace submodule + +/-! ### Submodules are cones -/ + +section ordered_semiring +variables [ordered_semiring 𝕜] + +section add_comm_monoid +variables [add_comm_monoid E] [module 𝕜 E] + +/-- Every submodule is trivially a convex cone. -/ +def to_convex_cone (S : submodule 𝕜 E) : convex_cone 𝕜 E := +{ carrier := S, + smul_mem' := λ c hc x hx, S.smul_mem c hx, + add_mem' := λ x hx y hy, S.add_mem hx hy } + +@[simp] lemma coe_to_convex_cone (S : submodule 𝕜 E) : ↑S.to_convex_cone = (S : set E) := rfl + +@[simp] lemma mem_to_convex_cone {x : E} {S : submodule 𝕜 E} : x ∈ S.to_convex_cone ↔ x ∈ S := +iff.rfl + +@[simp] lemma to_convex_cone_le_iff {S T : submodule 𝕜 E} : + S.to_convex_cone ≤ T.to_convex_cone ↔ S ≤ T := +iff.rfl + +@[simp] lemma to_convex_cone_bot : (⊥ : submodule 𝕜 E).to_convex_cone = 0 := rfl +@[simp] lemma to_convex_cone_top : (⊤ : submodule 𝕜 E).to_convex_cone = ⊤ := rfl + +@[simp] lemma to_convex_cone_inf (S T : submodule 𝕜 E) : + (S ⊓ T).to_convex_cone = S.to_convex_cone ⊓ T.to_convex_cone := +rfl + +@[simp] lemma pointed_to_convex_cone (S : submodule 𝕜 E) : S.to_convex_cone.pointed := S.zero_mem + +end add_comm_monoid +end ordered_semiring + +end submodule + +namespace convex_cone + /-! ### Positive cone of an ordered module -/ section positive_cone From 85158adc774e81d5432d37a16a7fa9dd8cb5f596 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 12 Dec 2022 09:04:20 +0000 Subject: [PATCH 0015/1291] fix(analysis/normed_space/basic): add back computable module instance (#17813) This was added in #8164 but removed in #17804. This also removes `noncomputable theory` since in fact only one result in the file needed it! --- src/analysis/normed_space/basic.lean | 7 +++++-- 1 file changed, 5 insertions(+), 2 deletions(-) diff --git a/src/analysis/normed_space/basic.lean b/src/analysis/normed_space/basic.lean index 598b8af33e5a5..1d493c820861b 100644 --- a/src/analysis/normed_space/basic.lean +++ b/src/analysis/normed_space/basic.lean @@ -17,7 +17,6 @@ about these definitions. variables {α : Type*} {β : Type*} {γ : Type*} {ι : Type*} -noncomputable theory open filter metric function set open_locale topological_space big_operators nnreal ennreal uniformity pointwise @@ -49,6 +48,10 @@ instance normed_space.has_bounded_smul [normed_space α β] : has_bounded_smul dist_pair_smul' := λ x₁ x₂ y, by simpa [dist_eq_norm, sub_smul] using normed_space.norm_smul_le (x₁ - x₂) y } +-- Shortcut instance, as otherwise this will be found by `normed_space.to_module` and be +-- noncomputable. +instance : module ℝ ℝ := by apply_instance + instance normed_field.to_normed_space : normed_space α α := { norm_smul_le := λ a b, le_of_eq (norm_mul a b) } @@ -186,7 +189,7 @@ In many cases the actual implementation is not important, so we don't mark the p See also `cont_diff_homeomorph_unit_ball` and `cont_diff_on_homeomorph_unit_ball_symm` for smoothness properties that hold when `E` is an inner-product space. -/ @[simps { attrs := [] }] -def homeomorph_unit_ball [normed_space ℝ E] : +noncomputable def homeomorph_unit_ball [normed_space ℝ E] : E ≃ₜ ball (0 : E) 1 := { to_fun := λ x, ⟨(1 + ‖x‖^2).sqrt⁻¹ • x, begin have : 0 < 1 + ‖x‖ ^ 2, by positivity, From 6d584f1709bedbed9175bd9350df46599bdd7213 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 12 Dec 2022 09:04:21 +0000 Subject: [PATCH 0016/1291] refactor(linear_algebra/basic): use `bijective` in `linear_equiv.of_bijective` (#17897) The two-argument version is a remnant of when we used to express these in terms of `range` and `ker`, which we dropped due to wanting to support semirings. We already have `equiv.of_bijective`, `ring_equiv.of_bijective`, etc which take a single `function.bijective` argument instead of two arguments; this makes the `linear_equiv` and `lie_equiv` spellings consistent. --- src/algebra/category/Module/subobject.lean | 1 + src/algebra/direct_sum/module.lean | 2 +- src/algebra/lie/basic.lean | 4 ++-- src/algebra/lie/subalgebra.lean | 4 ++-- src/algebra/module/pid.lean | 2 +- src/analysis/inner_product_space/pi_L2.lean | 4 ++-- src/analysis/normed_space/banach.lean | 2 +- src/analysis/normed_space/linear_isometry.lean | 2 +- src/linear_algebra/basic.lean | 6 +++--- src/linear_algebra/dimension.lean | 2 +- src/linear_algebra/dual.lean | 6 +++--- src/linear_algebra/finite_dimensional.lean | 6 +++--- src/linear_algebra/isomorphisms.lean | 6 +++--- src/linear_algebra/linear_independent.lean | 1 + src/linear_algebra/projection.lean | 9 +++++---- src/number_theory/ramification_inertia.lean | 2 +- src/ring_theory/is_tensor_product.lean | 2 +- src/ring_theory/witt_vector/isocrystal.lean | 2 +- src/topology/algebra/module/finite_dimension.lean | 2 +- 19 files changed, 34 insertions(+), 31 deletions(-) diff --git a/src/algebra/category/Module/subobject.lean b/src/algebra/category/Module/subobject.lean index 5455ae93abc23..ea03103006f3c 100644 --- a/src/algebra/category/Module/subobject.lean +++ b/src/algebra/category/Module/subobject.lean @@ -37,6 +37,7 @@ noncomputable def subobject_Module : subobject M ≃o submodule R M := order_iso fapply eq_mk_of_comm, { apply linear_equiv.to_Module_iso'_left, apply linear_equiv.of_bijective (linear_map.cod_restrict S.arrow.range S.arrow _), + split, { simpa only [← linear_map.ker_eq_bot, linear_map.ker_cod_restrict] using ker_eq_bot_of_mono _ }, { rw [← linear_map.range_eq_top, linear_map.range_cod_restrict, diff --git a/src/algebra/direct_sum/module.lean b/src/algebra/direct_sum/module.lean index cc94ec98c350e..a6c323a53a35f 100644 --- a/src/algebra/direct_sum/module.lean +++ b/src/algebra/direct_sum/module.lean @@ -286,7 +286,7 @@ noncomputable def is_internal.collected_basis (h : is_internal A) {α : ι → Type*} (v : Π i, basis (α i) R (A i)) : basis (Σ i, α i) R M := { repr := - (linear_equiv.of_bijective (direct_sum.coe_linear_map A) h.injective h.surjective).symm ≪≫ₗ + (linear_equiv.of_bijective (direct_sum.coe_linear_map A) h).symm ≪≫ₗ (dfinsupp.map_range.linear_equiv (λ i, (v i).repr)) ≪≫ₗ (sigma_finsupp_lequiv_dfinsupp R).symm } diff --git a/src/algebra/lie/basic.lean b/src/algebra/lie/basic.lean index 2fad21eb8ee7e..c0ca6f4394667 100644 --- a/src/algebra/lie/basic.lean +++ b/src/algebra/lie/basic.lean @@ -450,10 +450,10 @@ e.to_linear_equiv.surjective /-- A bijective morphism of Lie algebras yields an equivalence of Lie algebras. -/ @[simps] noncomputable def of_bijective (f : L₁ →ₗ⁅R⁆ L₂) - (h₁ : function.injective f) (h₂ : function.surjective f) : L₁ ≃ₗ⁅R⁆ L₂ := + (h : function.bijective f) : L₁ ≃ₗ⁅R⁆ L₂ := { to_fun := f, map_lie' := f.map_lie, - .. (linear_equiv.of_bijective (f : L₁ →ₗ[R] L₂) h₁ h₂), } + .. (linear_equiv.of_bijective (f : L₁ →ₗ[R] L₂) h), } end lie_equiv diff --git a/src/algebra/lie/subalgebra.lean b/src/algebra/lie/subalgebra.lean index b994a6914b853..db9b94416b120 100644 --- a/src/algebra/lie/subalgebra.lean +++ b/src/algebra/lie/subalgebra.lean @@ -242,11 +242,11 @@ end /-- A Lie algebra is equivalent to its range under an injective Lie algebra morphism. -/ noncomputable def equiv_range_of_injective (h : function.injective f) : L ≃ₗ⁅R⁆ f.range := -lie_equiv.of_bijective f.range_restrict (λ x y hxy, +lie_equiv.of_bijective f.range_restrict ⟨λ x y hxy, begin simp only [subtype.mk_eq_mk, range_restrict_apply] at hxy, exact h hxy, -end) f.surjective_range_restrict +end, f.surjective_range_restrict⟩ @[simp] lemma equiv_range_of_injective_apply (h : function.injective f) (x : L) : f.equiv_range_of_injective h x = ⟨f x, mem_range_self f x⟩ := diff --git a/src/algebra/module/pid.lean b/src/algebra/module/pid.lean index 4e5ee7755ede0..8c50e5b3ace4f 100644 --- a/src/algebra/module/pid.lean +++ b/src/algebra/module/pid.lean @@ -215,7 +215,7 @@ begin ((is_torsion'_powers_iff $ p i).mpr $ λ x, ⟨e i, smul_torsion_by _ _⟩) }, refine ⟨Σ i, fin (this i).some, infer_instance, λ ⟨i, j⟩, p i, λ ⟨i, j⟩, hp i, λ ⟨i, j⟩, (this i).some_spec.some j, - ⟨(linear_equiv.of_bijective (direct_sum.coe_linear_map _) h.1 h.2).symm.trans $ + ⟨(linear_equiv.of_bijective (direct_sum.coe_linear_map _) h).symm.trans $ (dfinsupp.map_range.linear_equiv $ λ i, (this i).some_spec.some_spec.some).trans $ (direct_sum.sigma_lcurry_equiv R).symm.trans (dfinsupp.map_range.linear_equiv $ λ i, quot_equiv_of_eq _ _ _)⟩⟩, diff --git a/src/analysis/inner_product_space/pi_L2.lean b/src/analysis/inner_product_space/pi_L2.lean index 1ed855d2f4947..e66ef1bc6a3c0 100644 --- a/src/analysis/inner_product_space/pi_L2.lean +++ b/src/analysis/inner_product_space/pi_L2.lean @@ -146,7 +146,7 @@ def direct_sum.is_internal.isometry_L2_of_orthogonal_family E ≃ₗᵢ[𝕜] pi_Lp 2 (λ i, V i) := begin let e₁ := direct_sum.linear_equiv_fun_on_fintype 𝕜 ι (λ i, V i), - let e₂ := linear_equiv.of_bijective (direct_sum.coe_linear_map V) hV.injective hV.surjective, + let e₂ := linear_equiv.of_bijective (direct_sum.coe_linear_map V) hV, refine (e₂.symm.trans e₁).isometry_of_inner _, suffices : ∀ v w, ⟪v, w⟫ = ⟪e₂ (e₁.symm v), e₂ (e₁.symm w)⟫, { intros v₀ w₀, @@ -166,7 +166,7 @@ end begin classical, let e₁ := direct_sum.linear_equiv_fun_on_fintype 𝕜 ι (λ i, V i), - let e₂ := linear_equiv.of_bijective (direct_sum.coe_linear_map V) hV.injective hV.surjective, + let e₂ := linear_equiv.of_bijective (direct_sum.coe_linear_map V) hV, suffices : ∀ v : ⨁ i, V i, e₂ v = ∑ i, e₁ v i, { exact this (e₁.symm w) }, intros v, diff --git a/src/analysis/normed_space/banach.lean b/src/analysis/normed_space/banach.lean index d086f173af75a..e34f631b63f55 100644 --- a/src/analysis/normed_space/banach.lean +++ b/src/analysis/normed_space/banach.lean @@ -336,7 +336,7 @@ variables [complete_space E] to a continuous linear equivalence. -/ noncomputable def of_bijective (f : E →L[𝕜] F) (hinj : ker f = ⊥) (hsurj : linear_map.range f = ⊤) : E ≃L[𝕜] F := -(linear_equiv.of_bijective ↑f (linear_map.ker_eq_bot.mp hinj) (linear_map.range_eq_top.mp hsurj)) +(linear_equiv.of_bijective ↑f ⟨linear_map.ker_eq_bot.mp hinj, linear_map.range_eq_top.mp hsurj⟩) .to_continuous_linear_equiv_of_continuous f.continuous @[simp] lemma coe_fn_of_bijective (f : E →L[𝕜] F) (hinj : ker f = ⊥) diff --git a/src/analysis/normed_space/linear_isometry.lean b/src/analysis/normed_space/linear_isometry.lean index 197d40bd68a70..f23d83eb76925 100644 --- a/src/analysis/normed_space/linear_isometry.lean +++ b/src/analysis/normed_space/linear_isometry.lean @@ -720,7 +720,7 @@ noncomputable def of_surjective (f : F →ₛₗᵢ[σ₁₂] E₂) (hfr : function.surjective f) : F ≃ₛₗᵢ[σ₁₂] E₂ := { norm_map' := f.norm_map, - .. linear_equiv.of_bijective f.to_linear_map f.injective hfr } + .. linear_equiv.of_bijective f.to_linear_map ⟨f.injective, hfr⟩ } @[simp] lemma coe_of_surjective (f : F →ₛₗᵢ[σ₁₂] E₂) (hfr : function.surjective f) : ⇑(linear_isometry_equiv.of_surjective f hfr) = f := diff --git a/src/linear_algebra/basic.lean b/src/linear_algebra/basic.lean index e8a0ed922c08a..5dec3792dc113 100644 --- a/src/linear_algebra/basic.lean +++ b/src/linear_algebra/basic.lean @@ -1778,11 +1778,11 @@ of_left_inverse $ classical.some_spec h.has_left_inverse /-- A bijective linear map is a linear equivalence. -/ noncomputable def of_bijective [ring_hom_inv_pair σ₁₂ σ₂₁] [ring_hom_inv_pair σ₂₁ σ₁₂] - (hf₁ : injective f) (hf₂ : surjective f) : M ≃ₛₗ[σ₁₂] M₂ := -(of_injective f hf₁).trans (of_top _ $ linear_map.range_eq_top.2 hf₂) + (hf : bijective f) : M ≃ₛₗ[σ₁₂] M₂ := +(of_injective f hf.injective).trans (of_top _ $ linear_map.range_eq_top.2 hf.surjective) @[simp] theorem of_bijective_apply [ring_hom_inv_pair σ₁₂ σ₂₁] [ring_hom_inv_pair σ₂₁ σ₁₂] - {hf₁ hf₂} (x : M) : of_bijective f hf₁ hf₂ x = f x := rfl + {hf} (x : M) : of_bijective f hf x = f x := rfl end diff --git a/src/linear_algebra/dimension.lean b/src/linear_algebra/dimension.lean index 072c7e7f8a19e..b3710244f1164 100644 --- a/src/linear_algebra/dimension.lean +++ b/src/linear_algebra/dimension.lean @@ -1056,7 +1056,7 @@ begin conv {to_rhs, rw [← dim_prod, dim_eq_of_surjective _ hf] }, congr' 1, apply linear_equiv.dim_eq, - refine linear_equiv.of_bijective _ _ _, + refine linear_equiv.of_bijective _ ⟨_, _⟩, { refine cod_restrict _ (prod cd (- ce)) _, { assume c, simp only [add_eq_zero_iff_eq_neg, linear_map.prod_apply, mem_ker, pi.prod, diff --git a/src/linear_algebra/dual.lean b/src/linear_algebra/dual.lean index ddd7a4282982c..98e9fc9892c8e 100644 --- a/src/linear_algebra/dual.lean +++ b/src/linear_algebra/dual.lean @@ -210,7 +210,7 @@ variables [_root_.finite ι] @[simps] def to_dual_equiv : M ≃ₗ[R] dual R M := linear_equiv.of_bijective b.to_dual - (ker_eq_bot.mp b.to_dual_ker) (range_eq_top.mp b.to_dual_range) + ⟨ker_eq_bot.mp b.to_dual_ker, range_eq_top.mp b.to_dual_range⟩ /-- Maps a basis for `V` to a basis for the dual space. -/ def dual_basis : basis ι R (dual R M) := b.map b.to_dual_equiv @@ -269,7 +269,7 @@ end /-- A module with a basis is linearly equivalent to the dual of its dual space. -/ def eval_equiv {ι : Type*} [_root_.finite ι] (b : basis ι R M) : M ≃ₗ[R] dual R (dual R M) := linear_equiv.of_bijective (eval R M) - (ker_eq_bot.mp b.eval_ker) (range_eq_top.mp b.eval_range) + ⟨ker_eq_bot.mp b.eval_ker, range_eq_top.mp b.eval_range⟩ @[simp] lemma eval_equiv_to_linear_map {ι : Type*} [_root_.finite ι] (b : basis ι R M) : (b.eval_equiv).to_linear_map = dual.eval R M := rfl @@ -347,7 +347,7 @@ variables (K V) /-- A vector space is linearly equivalent to the dual of its dual space. -/ def eval_equiv [finite_dimensional K V] : V ≃ₗ[K] dual K (dual K V) := linear_equiv.of_bijective (eval K V) - (ker_eq_bot.mp eval_ker) (range_eq_top.mp erange_coe) + ⟨ker_eq_bot.mp eval_ker, range_eq_top.mp erange_coe⟩ variables {K V} diff --git a/src/linear_algebra/finite_dimensional.lean b/src/linear_algebra/finite_dimensional.lean index a8dc37a15b4b3..1b98f9bb85e4c 100644 --- a/src/linear_algebra/finite_dimensional.lean +++ b/src/linear_algebra/finite_dimensional.lean @@ -943,7 +943,7 @@ variables [finite_dimensional K V] /-- The linear equivalence corresponging to an injective endomorphism. -/ noncomputable def of_injective_endo (f : V →ₗ[K] V) (h_inj : injective f) : V ≃ₗ[K] V := -linear_equiv.of_bijective f h_inj $ linear_map.injective_iff_surjective.mp h_inj +linear_equiv.of_bijective f ⟨h_inj, linear_map.injective_iff_surjective.mp h_inj⟩ @[simp] lemma coe_of_injective_endo (f : V →ₗ[K] V) (h_inj : injective f) : ⇑(of_injective_endo f h_inj) = f := rfl @@ -1031,8 +1031,8 @@ between the two vector spaces. -/ noncomputable def linear_equiv_of_injective [finite_dimensional K V] [finite_dimensional K V₂] (f : V →ₗ[K] V₂) (hf : injective f) (hdim : finrank K V = finrank K V₂) : V ≃ₗ[K] V₂ := -linear_equiv.of_bijective f hf $ - (linear_map.injective_iff_surjective_of_finrank_eq_finrank hdim).mp hf +linear_equiv.of_bijective f ⟨hf, + (linear_map.injective_iff_surjective_of_finrank_eq_finrank hdim).mp hf⟩ @[simp] lemma linear_equiv_of_injective_apply [finite_dimensional K V] [finite_dimensional K V₂] diff --git a/src/linear_algebra/isomorphisms.lean b/src/linear_algebra/isomorphisms.lean index a362f4542c7af..b77b3abdc038b 100644 --- a/src/linear_algebra/isomorphisms.lean +++ b/src/linear_algebra/isomorphisms.lean @@ -66,18 +66,18 @@ Second Isomorphism Law : the canonical map from `p/(p ∩ p')` to `(p+p')/p'` as noncomputable def quotient_inf_equiv_sup_quotient (p p' : submodule R M) : (p ⧸ (comap p.subtype (p ⊓ p'))) ≃ₗ[R] _ ⧸ (comap (p ⊔ p').subtype p') := by exact linear_equiv.of_bijective (quotient_inf_to_sup_quotient p p') - begin + ⟨begin rw [← ker_eq_bot, quotient_inf_to_sup_quotient, ker_liftq_eq_bot], rw [ker_comp, ker_mkq], exact λ ⟨x, hx1⟩ hx2, ⟨hx1, hx2⟩ - end + end, begin rw [← range_eq_top, quotient_inf_to_sup_quotient, range_liftq, eq_top_iff'], rintros ⟨x, hx⟩, rcases mem_sup.1 hx with ⟨y, hy, z, hz, rfl⟩, use [⟨y, hy⟩], apply (submodule.quotient.eq _).2, change y - (y + z) ∈ p', rwa [sub_add_eq_sub_sub, sub_self, zero_sub, neg_mem_iff] - end + end⟩ @[simp] lemma coe_quotient_inf_to_sup_quotient (p p' : submodule R M) : ⇑(quotient_inf_to_sup_quotient p p') = quotient_inf_equiv_sup_quotient p p' := rfl diff --git a/src/linear_algebra/linear_independent.lean b/src/linear_algebra/linear_independent.lean index b0baca789a72f..ce81c6baf68e4 100644 --- a/src/linear_algebra/linear_independent.lean +++ b/src/linear_algebra/linear_independent.lean @@ -703,6 +703,7 @@ def linear_independent.total_equiv (hv : linear_independent R v) : begin apply linear_equiv.of_bijective (linear_map.cod_restrict (span R (range v)) (finsupp.total ι M R v) _), + split, { rw [← linear_map.ker_eq_bot, linear_map.ker_cod_restrict], apply hv }, { rw [← linear_map.range_eq_top, linear_map.range_eq_map, linear_map.map_cod_restrict, diff --git a/src/linear_algebra/projection.lean b/src/linear_algebra/projection.lean index ac662dac37e42..0607c59f7ad11 100644 --- a/src/linear_algebra/projection.lean +++ b/src/linear_algebra/projection.lean @@ -75,8 +75,8 @@ open linear_map /-- If `q` is a complement of `p`, then `M/p ≃ q`. -/ def quotient_equiv_of_is_compl (h : is_compl p q) : (E ⧸ p) ≃ₗ[R] q := linear_equiv.symm $ linear_equiv.of_bijective (p.mkq.comp q.subtype) - (by rw [← ker_eq_bot, ker_comp, ker_mkq, disjoint_iff_comap_eq_bot.1 h.symm.disjoint]) - (by rw [← range_eq_top, range_comp, range_subtype, map_mkq_eq_top, h.sup_eq_top]) + ⟨by rw [← ker_eq_bot, ker_comp, ker_mkq, disjoint_iff_comap_eq_bot.1 h.symm.disjoint], + by rw [← range_eq_top, range_comp, range_subtype, map_mkq_eq_top, h.sup_eq_top]⟩ @[simp] lemma quotient_equiv_of_is_compl_symm_apply (h : is_compl p q) (x : q) : (quotient_equiv_of_is_compl p q h).symm x = quotient.mk x := rfl @@ -94,6 +94,7 @@ linear map `f : E → p` such that `f x = x` for `x ∈ p` and `f x = 0` for `x def prod_equiv_of_is_compl (h : is_compl p q) : (p × q) ≃ₗ[R] E := begin apply linear_equiv.of_bijective (p.subtype.coprod q.subtype), + split, { rw [← ker_eq_bot, ker_coprod_of_disjoint_range, ker_subtype, ker_subtype, prod_bot], rw [range_subtype, range_subtype], exact h.1 }, @@ -297,8 +298,8 @@ a linear equivalence `E ≃ₗ[R] F × G`. -/ def equiv_prod_of_surjective_of_is_compl (f : E →ₗ[R] F) (g : E →ₗ[R] G) (hf : f.range = ⊤) (hg : g.range = ⊤) (hfg : is_compl f.ker g.ker) : E ≃ₗ[R] F × G := -linear_equiv.of_bijective (f.prod g) (by simp [← ker_eq_bot, hfg.inf_eq_bot]) - (by { rw [←range_eq_top], simp [range_prod_eq hfg.sup_eq_top, *] }) +linear_equiv.of_bijective (f.prod g) ⟨by simp [← ker_eq_bot, hfg.inf_eq_bot], + by { rw [←range_eq_top], simp [range_prod_eq hfg.sup_eq_top, *] }⟩ @[simp] lemma coe_equiv_prod_of_surjective_of_is_compl {f : E →ₗ[R] F} {g : E →ₗ[R] G} (hf : f.range = ⊤) (hg : g.range = ⊤) (hfg : is_compl f.ker g.ker) : diff --git a/src/number_theory/ramification_inertia.lean b/src/number_theory/ramification_inertia.lean index cd02289a4ccc1..f957423607628 100644 --- a/src/number_theory/ramification_inertia.lean +++ b/src/number_theory/ramification_inertia.lean @@ -598,7 +598,7 @@ noncomputable def quotient_range_pow_quot_succ_inclusion_equiv [is_domain S] [is begin choose a a_mem a_not_mem using set_like.exists_of_lt (ideal.strict_anti_pow P hP (ideal.is_prime.ne_top infer_instance) (le_refl i.succ)), - refine (linear_equiv.of_bijective _ _ _).symm, + refine (linear_equiv.of_bijective _ ⟨_, _⟩).symm, { exact quotient_to_quotient_range_pow_quot_succ f p P a_mem }, { exact quotient_to_quotient_range_pow_quot_succ_injective f p P hi a_mem a_not_mem }, { exact quotient_to_quotient_range_pow_quot_succ_surjective f p P hP hi a_mem a_not_mem } diff --git a/src/ring_theory/is_tensor_product.lean b/src/ring_theory/is_tensor_product.lean index b939d0ea3c02a..50b6eac2fbaef 100644 --- a/src/ring_theory/is_tensor_product.lean +++ b/src/ring_theory/is_tensor_product.lean @@ -69,7 +69,7 @@ variables {R M N} /-- If `M` is the tensor product of `M₁` and `M₂`, it is linearly equivalent to `M₁ ⊗[R] M₂`. -/ @[simps apply] noncomputable def is_tensor_product.equiv (h : is_tensor_product f) : M₁ ⊗[R] M₂ ≃ₗ[R] M := -linear_equiv.of_bijective _ h.1 h.2 +linear_equiv.of_bijective _ h @[simp] lemma is_tensor_product.equiv_to_linear_map (h : is_tensor_product f) : h.equiv.to_linear_map = tensor_product.lift f := rfl diff --git a/src/ring_theory/witt_vector/isocrystal.lean b/src/ring_theory/witt_vector/isocrystal.lean index 74f49caeefa05..bc06cef87421a 100644 --- a/src/ring_theory/witt_vector/isocrystal.lean +++ b/src/ring_theory/witt_vector/isocrystal.lean @@ -183,7 +183,7 @@ begin let F₀ : standard_one_dim_isocrystal p k m →ₗ[K(p,k)] V := linear_map.to_span_singleton K(p, k) V x, let F : standard_one_dim_isocrystal p k m ≃ₗ[K(p,k)] V, - { refine linear_equiv.of_bijective F₀ _ _, + { refine linear_equiv.of_bijective F₀ ⟨_, _⟩, { rw ← linear_map.ker_eq_bot, exact linear_map.ker_to_span_singleton K(p, k) V hx }, { rw ← linear_map.range_eq_top, diff --git a/src/topology/algebra/module/finite_dimension.lean b/src/topology/algebra/module/finite_dimension.lean index c1d55c8cbcf35..fece5d836a471 100644 --- a/src/topology/algebra/module/finite_dimension.lean +++ b/src/topology/algebra/module/finite_dimension.lean @@ -168,7 +168,7 @@ begin have hs : function.surjective (l.ker.liftq l (le_refl _)), { rw [← linear_map.range_eq_top, submodule.range_liftq], exact eq_top_of_finrank_eq ((finrank_self 𝕜).symm ▸ this) }, - let φ : (E ⧸ l.ker) ≃ₗ[𝕜] 𝕜 := linear_equiv.of_bijective (l.ker.liftq l (le_refl _)) hi hs, + let φ : (E ⧸ l.ker) ≃ₗ[𝕜] 𝕜 := linear_equiv.of_bijective (l.ker.liftq l (le_refl _)) ⟨hi, hs⟩, have hlφ : (l : E → 𝕜) = φ ∘ l.ker.mkq, by ext; refl, -- Since the quotient map `E →ₗ[𝕜] (E ⧸ l.ker)` is continuous, the continuity of `l` will follow From bd59f822dd31da0b6eb683edd04ec01311210066 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 12 Dec 2022 09:04:22 +0000 Subject: [PATCH 0017/1291] feat(data/set/basic): A set is either a subsingleton or nontrivial (#17901) Also make the argument to `set.finite_or_infinite` explicit. --- src/data/finset/basic.lean | 5 ++++- src/data/set/basic.lean | 6 ++++++ src/data/set/finite.lean | 4 ++-- src/order/well_founded_set.lean | 2 +- 4 files changed, 13 insertions(+), 4 deletions(-) diff --git a/src/data/finset/basic.lean b/src/data/finset/basic.lean index c06c20d3f9f90..2c020aac5b057 100644 --- a/src/data/finset/basic.lean +++ b/src/data/finset/basic.lean @@ -432,7 +432,7 @@ end empty /-! ### singleton -/ section singleton -variables {a : α} +variables {s : finset α} {a : α} /-- `{a} : finset a` is the set `{a}` containing `a` and nothing else. @@ -520,6 +520,9 @@ by rw [←coe_ssubset, coe_singleton, set.ssubset_singleton_iff, coe_eq_empty] lemma eq_empty_of_ssubset_singleton {s : finset α} {x : α} (hs : s ⊂ {x}) : s = ∅ := ssubset_singleton_iff.1 hs +lemma eq_singleton_or_nontrivial (ha : a ∈ s) : s = {a} ∨ (s : set α).nontrivial := +by { rw ←coe_eq_singleton, exact set.eq_singleton_or_nontrivial ha } + instance [nonempty α] : nontrivial (finset α) := ‹nonempty α›.elim $ λ a, ⟨⟨{a}, ∅, singleton_ne_empty _⟩⟩ diff --git a/src/data/set/basic.lean b/src/data/set/basic.lean index 6dbbb57a6e255..24571fb1e9f51 100644 --- a/src/data/set/basic.lean +++ b/src/data/set/basic.lean @@ -1549,6 +1549,12 @@ iff.not_left not_subsingleton_iff.symm alias not_nontrivial_iff ↔ _ subsingleton.not_nontrivial alias not_subsingleton_iff ↔ _ nontrivial.not_subsingleton +protected lemma subsingleton_or_nontrivial (s : set α) : s.subsingleton ∨ s.nontrivial := +by simp [or_iff_not_imp_right] + +lemma eq_singleton_or_nontrivial (ha : a ∈ s) : s = {a} ∨ s.nontrivial := +by { rw ←subsingleton_iff_singleton ha, exact s.subsingleton_or_nontrivial } + theorem univ_eq_true_false : univ = ({true, false} : set Prop) := eq.symm $ eq_univ_of_forall $ classical.cases (by simp) (by simp) diff --git a/src/data/set/finite.lean b/src/data/set/finite.lean index d52c32394fa31..33841e128ac1c 100644 --- a/src/data/set/finite.lean +++ b/src/data/set/finite.lean @@ -114,8 +114,8 @@ protected def infinite (s : set α) : Prop := ¬ s.finite @[simp] lemma not_infinite {s : set α} : ¬ s.infinite ↔ s.finite := not_not -/-- See also `fintype_or_infinite`. -/ -lemma finite_or_infinite {s : set α} : s.finite ∨ s.infinite := em _ +/-- See also `finite_or_infinite`, `fintype_or_infinite`. -/ +protected lemma finite_or_infinite (s : set α) : s.finite ∨ s.infinite := em _ /-! ### Basic properties of `set.finite.to_finset` -/ diff --git a/src/order/well_founded_set.lean b/src/order/well_founded_set.lean index 4fa757296032f..f726237ba0da4 100644 --- a/src/order/well_founded_set.lean +++ b/src/order/well_founded_set.lean @@ -251,7 +251,7 @@ lemma _root_.is_antichain.finite_of_partially_well_ordered_on (ha : is_antichain (hp : s.partially_well_ordered_on r) : s.finite := begin - refine finite_or_infinite.resolve_right (λ hi, _), + refine not_infinite.1 (λ hi, _), obtain ⟨m, n, hmn, h⟩ := hp (λ n, hi.nat_embedding _ n) (λ n, (hi.nat_embedding _ n).2), exact hmn.ne ((hi.nat_embedding _).injective $ subtype.val_injective $ ha.eq (hi.nat_embedding _ m).2 (hi.nat_embedding _ n).2 h), From ee00ca5e2fb606ac5fcb6a48245c3eb2f458b1b5 Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Mon, 12 Dec 2022 12:35:59 +0000 Subject: [PATCH 0018/1291] feat(linear_algebra/affine_space/independent): spans of two affine combinations (#17861) Add some lemmas about spans of two points expressed as affine combinations. --- .../affine_space/independent.lean | 102 ++++++++++++++++++ 1 file changed, 102 insertions(+) diff --git a/src/linear_algebra/affine_space/independent.lean b/src/linear_algebra/affine_space/independent.lean index 4d42768003044..2e485e13f0703 100644 --- a/src/linear_algebra/affine_space/independent.lean +++ b/src/linear_algebra/affine_space/independent.lean @@ -5,6 +5,7 @@ Authors: Joseph Myers -/ import data.finset.sort import data.fin.vec_notation +import data.sign import linear_algebra.affine_space.combination import linear_algebra.affine_space.affine_equiv import linear_algebra.basis @@ -471,6 +472,55 @@ lemma affine_independent_iff {ι} {p : ι → V} : ∀ (s : finset ι) (w : ι → k), s.sum w = 0 → ∑ e in s, w e • p e = 0 → ∀ (e ∈ s), w e = 0 := forall₃_congr (λ s w hw, by simp [s.weighted_vsub_eq_linear_combination hw]) +/-- Given an affinely independent family of points, a weighted subtraction lies in the +`vector_span` of two points given as affine combinations if and only if it is a weighted +subtraction with weights a multiple of the difference between the weights of the two points. -/ +lemma weighted_vsub_mem_vector_span_pair {p : ι → P} (h : affine_independent k p) + {w w₁ w₂ : ι → k} {s : finset ι} (hw : ∑ i in s, w i = 0) (hw₁ : ∑ i in s, w₁ i = 1) + (hw₂ : ∑ i in s, w₂ i = 1) : + s.weighted_vsub p w ∈ + vector_span k ({s.affine_combination p w₁, s.affine_combination p w₂} : set P) ↔ + ∃ r : k, ∀ i ∈ s, w i = r * (w₁ i - w₂ i) := +begin + rw mem_vector_span_pair, + refine ⟨λ h, _, λ h, _⟩, + { rcases h with ⟨r, hr⟩, + refine ⟨r, λ i hi, _⟩, + rw [s.affine_combination_vsub, ←s.weighted_vsub_const_smul, ←sub_eq_zero, ←map_sub] at hr, + have hw' : ∑ j in s, (r • (w₁ - w₂) - w) j = 0, + { simp_rw [pi.sub_apply, pi.smul_apply, pi.sub_apply, smul_sub, finset.sum_sub_distrib, + ←finset.smul_sum, hw, hw₁, hw₂, sub_self] }, + have hr' := h s _ hw' hr i hi, + rw [eq_comm, ←sub_eq_zero, ←smul_eq_mul], + exact hr' }, + { rcases h with ⟨r, hr⟩, + refine ⟨r, _⟩, + let w' := λ i, r * (w₁ i - w₂ i), + change ∀ i ∈ s, w i = w' i at hr, + rw [s.weighted_vsub_congr hr (λ _ _, rfl), s.affine_combination_vsub, + ←s.weighted_vsub_const_smul], + congr } +end + +/-- Given an affinely independent family of points, an affine combination lies in the +span of two points given as affine combinations if and only if it is an affine combination +with weights those of one point plus a multiple of the difference between the weights of the +two points. -/ +lemma affine_combination_mem_affine_span_pair {p : ι → P} (h : affine_independent k p) + {w w₁ w₂ : ι → k} {s : finset ι} (hw : ∑ i in s, w i = 1) (hw₁ : ∑ i in s, w₁ i = 1) + (hw₂ : ∑ i in s, w₂ i = 1) : + s.affine_combination p w ∈ + line[k, s.affine_combination p w₁, s.affine_combination p w₂] ↔ + ∃ r : k, ∀ i ∈ s, w i = r * (w₂ i - w₁ i) + w₁ i := +begin + rw [←vsub_vadd (s.affine_combination p w) (s.affine_combination p w₁), + affine_subspace.vadd_mem_iff_mem_direction _ (left_mem_affine_span_pair _ _ _), + direction_affine_span, s.affine_combination_vsub, set.pair_comm, + weighted_vsub_mem_vector_span_pair h _ hw₂ hw₁], + { simp only [pi.sub_apply, sub_eq_iff_eq_add] }, + { simp_rw [pi.sub_apply, finset.sum_sub_distrib, hw, hw₁, sub_self] } +end + end affine_independent section division_ring @@ -653,6 +703,58 @@ end end division_ring +section ordered + +variables {k : Type*} {V : Type*} {P : Type*} [linear_ordered_ring k] [add_comm_group V] +variables [module k V] [affine_space V P] {ι : Type*} +include V + +local attribute [instance] linear_ordered_ring.decidable_lt + +/-- Given an affinely independent family of points, suppose that an affine combination lies in +the span of two points given as affine combinations, and suppose that, for two indices, the +coefficients in the first point in the span are zero and those in the second point in the span +have the same sign. Then the coefficients in the combination lying in the span have the same +sign. -/ +lemma sign_eq_of_affine_combination_mem_affine_span_pair {p : ι → P} (h : affine_independent k p) + {w w₁ w₂ : ι → k} {s : finset ι} (hw : ∑ i in s, w i = 1) (hw₁ : ∑ i in s, w₁ i = 1) + (hw₂ : ∑ i in s, w₂ i = 1) + (hs : s.affine_combination p w ∈ line[k, s.affine_combination p w₁, s.affine_combination p w₂]) + {i j : ι} (hi : i ∈ s) (hj : j ∈ s) (hi0 : w₁ i = 0) (hj0 : w₁ j = 0) + (hij : sign (w₂ i) = sign (w₂ j)) : sign (w i) = sign (w j) := +begin + rw affine_combination_mem_affine_span_pair h hw hw₁ hw₂ at hs, + rcases hs with ⟨r, hr⟩, + dsimp only at hr, + rw [hr i hi, hr j hj, hi0, hj0, add_zero, add_zero, sub_zero, sub_zero, sign_mul, sign_mul, hij] +end + +/-- Given an affinely independent family of points, suppose that an affine combination lies in +the span of one point of that family and a combination of another two points of that family given +by `line_map` with coefficient between 0 and 1. Then the coefficients of those two points in the +combination lying in the span have the same sign. -/ +lemma sign_eq_of_affine_combination_mem_affine_span_single_line_map {p : ι → P} + (h : affine_independent k p) {w : ι → k} {s : finset ι} (hw : ∑ i in s, w i = 1) + {i₁ i₂ i₃ : ι} (h₁ : i₁ ∈ s) (h₂ : i₂ ∈ s) (h₃ : i₃ ∈ s) (h₁₂ : i₁ ≠ i₂) (h₁₃ : i₁ ≠ i₃) + (h₂₃ : i₂ ≠ i₃) {c : k} (hc0 : 0 < c) (hc1 : c < 1) + (hs : s.affine_combination p w ∈ line[k, p i₁, affine_map.line_map (p i₂) (p i₃) c]) : + sign (w i₂) = sign (w i₃) := +begin + classical, + rw [←s.affine_combination_affine_combination_single_weights k p h₁, + ←s.affine_combination_affine_combination_line_map_weights p h₂ h₃ c] at hs, + refine sign_eq_of_affine_combination_mem_affine_span_pair h hw + (s.sum_affine_combination_single_weights k h₁) + (s.sum_affine_combination_line_map_weights h₂ h₃ c) hs h₂ h₃ + (finset.affine_combination_single_weights_apply_of_ne k h₁₂.symm) + (finset.affine_combination_single_weights_apply_of_ne k h₁₃.symm) _, + rw [finset.affine_combination_line_map_weights_apply_left h₂₃, + finset.affine_combination_line_map_weights_apply_right h₂₃], + simp [hc0, sub_pos.2 hc1] +end + +end ordered + namespace affine variables (k : Type*) {V : Type*} (P : Type*) [ring k] [add_comm_group V] [module k V] From 32d096dc3c4bb9df248d66874f0489eef1f70419 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck <58564076+AlexKontorovich@users.noreply.github.com> Date: Mon, 12 Dec 2022 14:47:32 +0000 Subject: [PATCH 0019/1291] feat(number_theory/modular_forms): Modular form definition (#13250) This contains the basic definitions of modular forms and cusp forms together with some basic lemmas. Co-authored-by: Johan Commelin --- .../functions_bounded_at_infty.lean | 6 +- src/number_theory/modular_forms/basic.lean | 299 ++++++++++++++++++ .../modular_forms/congruence_subgroups.lean | 10 +- .../modular_forms/slash_actions.lean | 41 +-- .../modular_forms/slash_invariant_forms.lean | 4 +- .../filter/zero_and_bounded_at_filter.lean | 82 +++-- 6 files changed, 397 insertions(+), 45 deletions(-) create mode 100644 src/number_theory/modular_forms/basic.lean diff --git a/src/analysis/complex/upper_half_plane/functions_bounded_at_infty.lean b/src/analysis/complex/upper_half_plane/functions_bounded_at_infty.lean index 7700766e8f17b..3a6ab5170dd87 100644 --- a/src/analysis/complex/upper_half_plane/functions_bounded_at_infty.lean +++ b/src/analysis/complex/upper_half_plane/functions_bounded_at_infty.lean @@ -42,7 +42,7 @@ begin end /-- A function ` f : ℍ → α` is bounded at infinity if it is bounded along `at_im_infty`. -/ -def is_bounded_at_im_infty {α : Type*} [has_norm α] [has_one (ℍ → α)] (f : ℍ → α) : Prop := +def is_bounded_at_im_infty {α : Type*} [has_norm α] (f : ℍ → α) : Prop := bounded_at_filter at_im_infty f /-- A function ` f : ℍ → α` is zero at infinity it is zero along `at_im_infty`. -/ @@ -50,7 +50,7 @@ def is_zero_at_im_infty {α : Type*} [has_zero α] [topological_space α] (f : zero_at_filter at_im_infty f lemma zero_form_is_bounded_at_im_infty {α : Type*} [normed_field α] : - is_bounded_at_im_infty (0 : ℍ → α) := zero_is_bounded_at_filter at_im_infty + is_bounded_at_im_infty (0 : ℍ → α) := const_bounded_at_filter at_im_infty (0:α) /-- Module of functions that are zero at infinity. -/ def zero_at_im_infty_submodule (α : Type*) [normed_field α] : submodule α (ℍ → α) := @@ -64,7 +64,7 @@ lemma is_bounded_at_im_infty.mul {f g : ℍ → ℂ} (hf : is_bounded_at_im_inft (hg : is_bounded_at_im_infty g) : is_bounded_at_im_infty (f * g) := by simpa only [pi.one_apply, mul_one, norm_eq_abs] using hf.mul hg -@[simp] lemma bounded_mem (f : ℍ → ℂ) : +lemma bounded_mem (f : ℍ → ℂ) : is_bounded_at_im_infty f ↔ ∃ (M A : ℝ), ∀ z : ℍ, A ≤ im z → abs (f z) ≤ M := by simp [is_bounded_at_im_infty, bounded_at_filter, asymptotics.is_O_iff, filter.eventually, at_im_infty_mem] diff --git a/src/number_theory/modular_forms/basic.lean b/src/number_theory/modular_forms/basic.lean new file mode 100644 index 0000000000000..d2811244e051a --- /dev/null +++ b/src/number_theory/modular_forms/basic.lean @@ -0,0 +1,299 @@ +/- +Copyright (c) 2022 Chris Birkbeck. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Chris Birkbeck +-/ + +import geometry.manifold.mfderiv +import analysis.complex.upper_half_plane.functions_bounded_at_infty +import analysis.complex.upper_half_plane.topology +import number_theory.modular_forms.slash_invariant_forms +/-! +# Modular forms + +This file defines modular forms and proves some basic properties about them. + +We begin by defining modular forms and cusp forms as extension of `slash_invariant_forms` then we +define the space of modular forms, cusp forms and prove that the product of two modular forms is a +modular form. +-/ + +open complex upper_half_plane + +open_locale topological_space manifold upper_half_plane + +noncomputable theory + +instance upper_half_plane.charted_space : charted_space ℂ ℍ := +upper_half_plane.open_embedding_coe.singleton_charted_space + +instance upper_half_plane.smooth_manifold_with_corners : smooth_manifold_with_corners 𝓘(ℂ) ℍ := +upper_half_plane.open_embedding_coe.singleton_smooth_manifold_with_corners 𝓘(ℂ) + +local prefix `↑ₘ`:1024 := @coe _ (matrix (fin 2) (fin 2) _) _ + +local notation `GL(` n `, ` R `)`⁺ := matrix.GL_pos (fin n) R + +local notation `SL(` n `, ` R `)` := matrix.special_linear_group (fin n) R + +section modular_form + +open modular_form + +variables (F : Type*) (Γ : subgroup SL(2, ℤ)) (k : ℤ) + +local notation f `∣[`:73 k:0, A `]` :72 := slash_action.map ℂ k A f + +set_option old_structure_cmd true + +/--These are `slash_invariant_form`'s that are holomophic and bounded at infinity. -/ +structure modular_form extends slash_invariant_form Γ k := +(holo' : mdifferentiable 𝓘(ℂ) 𝓘(ℂ) (to_fun : ℍ → ℂ)) +(bdd_at_infty' : ∀ (A : SL(2, ℤ)), is_bounded_at_im_infty (to_fun ∣[k, A])) + +/-- The `slash_invariant_form` associated to a `modular_form`. -/ +add_decl_doc modular_form.to_slash_invariant_form + +/--These are `slash_invariant_form`s that are holomophic and zero at infinity. -/ +structure cusp_form extends slash_invariant_form Γ k := +(holo' : mdifferentiable 𝓘(ℂ) 𝓘(ℂ) (to_fun : ℍ → ℂ)) +(zero_at_infty' : ∀ (A : SL(2, ℤ)), is_zero_at_im_infty (to_fun ∣[k, A])) + +/-- The `slash_invariant_form` associated to a `cusp_form`. -/ +add_decl_doc cusp_form.to_slash_invariant_form + +/--`modular_form_class F Γ k` says that `F` is a type of bundled functions that extend +`slash_invariant_forms_class` by requiring that the functions be holomorphic and bounded +at infinity. -/ +class modular_form_class extends slash_invariant_form_class F Γ k := +(holo: ∀ f : F, mdifferentiable 𝓘(ℂ) 𝓘(ℂ) (f : ℍ → ℂ)) +(bdd_at_infty : ∀ (f : F) (A : SL(2, ℤ)), is_bounded_at_im_infty (f ∣[k, A])) + +/--`cusp_form_class F Γ k` says that `F` is a type of bundled functions that extend +`slash_invariant_forms_class` by requiring that the functions be holomorphic and zero +at infinity. -/ +class cusp_form_class extends slash_invariant_form_class F Γ k := +(holo: ∀ f : F, mdifferentiable 𝓘(ℂ) 𝓘(ℂ) (f : ℍ → ℂ)) +(zero_at_infty : ∀ (f : F) (A : SL(2, ℤ)), is_zero_at_im_infty (f ∣[k, A])) + +@[priority 100] +instance modular_form_class.modular_form : modular_form_class (modular_form Γ k) Γ k := +{ coe := modular_form.to_fun, + coe_injective' := λ f g h, by cases f; cases g; congr', + slash_action_eq := modular_form.slash_action_eq', + holo:= modular_form.holo', + bdd_at_infty := modular_form.bdd_at_infty' } + +@[priority 100] +instance cusp_form_class.cusp_form : cusp_form_class (cusp_form Γ k) Γ k := +{ coe := cusp_form.to_fun, + coe_injective' := λ f g h, by cases f; cases g; congr', + slash_action_eq := cusp_form.slash_action_eq', + holo:= cusp_form.holo', + zero_at_infty := cusp_form.zero_at_infty' } + +variables {F Γ k} + +@[simp] lemma modular_form_to_fun_eq_coe {f : modular_form Γ k} : f.to_fun = (f : ℍ → ℂ) := rfl +@[simp] lemma cusp_form_to_fun_eq_coe {f : cusp_form Γ k} : f.to_fun = (f : ℍ → ℂ) := rfl + +@[ext] theorem modular_form.ext {f g : modular_form Γ k} (h : ∀ x, f x = g x) : f = g := +fun_like.ext f g h + +@[ext] theorem cusp_form.ext {f g : cusp_form Γ k} (h : ∀ x, f x = g x) : f = g := +fun_like.ext f g h + +/-- Copy of a `modular_form` with a new `to_fun` equal to the old one. Useful to fix +definitional equalities. -/ +protected def modular_form.copy (f : modular_form Γ k) (f' : ℍ → ℂ) (h : f' = ⇑f) : + modular_form Γ k := +{ to_fun := f', + slash_action_eq' := h.symm ▸ f.slash_action_eq', + holo' := h.symm ▸ f.holo', + bdd_at_infty' := λ A, h.symm ▸ f.bdd_at_infty' A } + +/-- Copy of a `cusp_form` with a new `to_fun` equal to the old one. Useful to fix +definitional equalities. -/ +protected def cusp_form.copy (f : cusp_form Γ k) (f' : ℍ → ℂ) (h : f' = ⇑f) : + cusp_form Γ k := +{ to_fun := f', + slash_action_eq' := h.symm ▸ f.slash_action_eq', + holo' := h.symm ▸ f.holo', + zero_at_infty' := λ A, h.symm ▸ f.zero_at_infty' A } + +end modular_form + +namespace modular_form + +open slash_invariant_form + +variables {F : Type*} {Γ : subgroup SL(2, ℤ)} {k : ℤ} + +instance has_add : has_add (modular_form Γ k) := +⟨ λ f g, + { holo' := f.holo'.add g.holo', + bdd_at_infty' := λ A, by simpa using (f.bdd_at_infty' A).add (g.bdd_at_infty' A), + .. (f : slash_invariant_form Γ k) + g }⟩ + +@[simp] lemma coe_add (f g : modular_form Γ k) : ⇑(f + g) = f + g := rfl + +@[simp] lemma add_apply (f g : modular_form Γ k) (z : ℍ) : (f + g) z = f z + g z := rfl + +instance has_zero : has_zero (modular_form Γ k) := +⟨ { holo' := (λ _, mdifferentiable_at_const 𝓘(ℂ, ℂ) 𝓘(ℂ, ℂ)), + bdd_at_infty' := λ A, by simpa using zero_form_is_bounded_at_im_infty, + .. (0 : slash_invariant_form Γ k) } ⟩ + +@[simp] lemma coe_zero : ⇑(0 : modular_form Γ k) = (0 : ℍ → ℂ) := rfl + +@[simp] lemma zero_apply (z : ℍ) : (0 : modular_form Γ k) z = 0 := rfl + +section +variables {α : Type*} [has_smul α ℂ] [is_scalar_tower α ℂ ℂ] + +instance has_smul : has_smul α (modular_form Γ k) := +⟨ λ c f, + { to_fun := c • f, + holo' := by simpa using f.holo'.const_smul (c • (1 : ℂ)), + bdd_at_infty' := λ A, by simpa using (f.bdd_at_infty' A).const_smul_left (c • (1 : ℂ)), + .. c • (f : slash_invariant_form Γ k)}⟩ + +@[simp] lemma coe_smul (f : (modular_form Γ k)) (n : α) : ⇑(n • f) = n • f := rfl +@[simp] lemma smul_apply (f : (modular_form Γ k)) (n : α) (z : ℍ) : + (n • f) z = n • (f z) := rfl +end + +instance has_neg : has_neg (modular_form Γ k) := +⟨λ f, + { to_fun := -f, + holo' := f.holo'.neg, + bdd_at_infty':= λ A, by simpa using (f.bdd_at_infty' A).neg, + .. -(f : slash_invariant_form Γ k) }⟩ + +@[simp] lemma coe_neg (f : modular_form Γ k) : ⇑(-f) = -f := rfl + +@[simp] lemma neg_apply (f : modular_form Γ k) (z : ℍ) : (-f) z = - (f z) := rfl + +instance has_sub : has_sub (modular_form Γ k) := +⟨ λ f g, f + -g ⟩ + +@[simp] lemma coe_sub (f g : (modular_form Γ k)) : ⇑(f - g) = f - g := rfl + +@[simp] lemma sub_apply (f g : modular_form Γ k) (z : ℍ) : (f - g) z = f z - g z := rfl + +instance : add_comm_group (modular_form Γ k) := +fun_like.coe_injective.add_comm_group _ rfl coe_add coe_neg coe_sub coe_smul coe_smul + +/--Additive coercion from `modular_form` to `ℍ → ℂ`. -/ +@[simps] def coe_hom : (modular_form Γ k) →+ (ℍ → ℂ) := +{ to_fun := λ f, f, + map_zero' := coe_zero, + map_add' := λ _ _, rfl } + +instance : module ℂ (modular_form Γ k) := +function.injective.module ℂ coe_hom fun_like.coe_injective (λ _ _, rfl) + +instance : inhabited (modular_form Γ k) := ⟨0⟩ + +/--The modular form of weight `k_1 + k_2` given by the product of two modular forms of weights +`k_1` and `k_2`. -/ +def mul {k_1 k_2 : ℤ} {Γ : subgroup SL(2, ℤ)} (f : (modular_form Γ k_1)) + (g : (modular_form Γ k_2)) : (modular_form Γ (k_1 + k_2)) := +{ to_fun := f * g, + slash_action_eq' := λ A, by simp_rw [mul_slash_subgroup, modular_form_class.slash_action_eq], + holo' := f.holo'.mul g.holo', + bdd_at_infty' := λ A, by simpa using (f.bdd_at_infty' A).mul (g.bdd_at_infty' A) } + +@[simp] lemma mul_coe {k_1 k_2 : ℤ} {Γ : subgroup SL(2, ℤ)} (f : (modular_form Γ k_1)) + (g : (modular_form Γ k_2)) : ((f.mul g) : ℍ → ℂ) = f * g := rfl + +instance : has_one (modular_form Γ 0) := +⟨{ holo' := λ x, mdifferentiable_at_const 𝓘(ℂ, ℂ) 𝓘(ℂ, ℂ), + bdd_at_infty' := λ A, by simpa using at_im_infty.const_bounded_at_filter (1:ℂ), + .. (1 : slash_invariant_form Γ 0) }⟩ + +@[simp] lemma one_coe_eq_one : ((1 : modular_form Γ 0) : ℍ → ℂ) = 1 := rfl + +end modular_form + +namespace cusp_form +open modular_form + +variables {F : Type*} {Γ : subgroup SL(2, ℤ)} {k : ℤ} + +instance has_add : has_add (cusp_form Γ k) := +⟨ λ f g, + { to_fun := f + g, + holo' := f.holo'.add g.holo', + zero_at_infty' := λ A, by simpa using (f.zero_at_infty' A).add (g.zero_at_infty' A), + .. (f : slash_invariant_form Γ k) + g }⟩ + +@[simp] lemma coe_add (f g : cusp_form Γ k) : ⇑(f + g) = f + g := rfl + +@[simp] lemma add_apply (f g : cusp_form Γ k) (z : ℍ) : (f + g) z = f z + g z := rfl + +instance has_zero : has_zero (cusp_form Γ k) := +⟨ { to_fun := 0, + holo' := (λ _, mdifferentiable_at_const 𝓘(ℂ, ℂ) 𝓘(ℂ, ℂ)), + zero_at_infty' := by simpa using filter.zero_zero_at_filter _, + .. (0 : slash_invariant_form Γ k) }⟩ + +@[simp] lemma coe_zero : ⇑(0 : cusp_form Γ k) = (0 : ℍ → ℂ) := rfl + +@[simp] lemma zero_apply (z : ℍ) : (0 : cusp_form Γ k) z = 0 := rfl + +section +variables {α : Type*} [has_smul α ℂ] [is_scalar_tower α ℂ ℂ] + +instance has_smul : has_smul α (cusp_form Γ k) := +⟨ λ c f, + { to_fun := c • f, + holo' := by simpa using f.holo'.const_smul (c • (1 : ℂ)), + zero_at_infty' := λ A, by simpa using (f.zero_at_infty' A).smul (c • (1 : ℂ)), + .. c • (f : slash_invariant_form Γ k) }⟩ + +@[simp] lemma coe_smul (f : (cusp_form Γ k)) (n : α) : ⇑(n • f) = n • f := rfl +@[simp] lemma smul_apply (f : (cusp_form Γ k)) (n : α) {z : ℍ} : + (n • f) z = n • (f z) := rfl + +end + +instance has_neg : has_neg (cusp_form Γ k) := +⟨λ f, + { to_fun := -f, + holo' := f.holo'.neg, + zero_at_infty':= λ A, by simpa using (f.zero_at_infty' A).neg, + .. -(f : slash_invariant_form Γ k)} ⟩ + +@[simp] lemma coe_neg (f : cusp_form Γ k) : ⇑(-f) = -f := rfl +@[simp] lemma neg_apply (f : cusp_form Γ k) (z : ℍ) : (-f) z = -(f z) := rfl + +instance has_sub : has_sub (cusp_form Γ k) := +⟨ λ f g, f + -g ⟩ + +@[simp] lemma coe_sub (f g : cusp_form Γ k) : ⇑(f - g) = f - g := rfl +@[simp] lemma sub_apply (f g : cusp_form Γ k) (z : ℍ) : (f - g) z = f z - g z := rfl + +instance : add_comm_group (cusp_form Γ k) := +fun_like.coe_injective.add_comm_group _ rfl coe_add coe_neg coe_sub coe_smul coe_smul + +/--Additive coercion from `cusp_form` to `ℍ → ℂ`. -/ +@[simps] def coe_hom : (cusp_form Γ k) →+ (ℍ → ℂ) := +{ to_fun := λ f, f, + map_zero' := cusp_form.coe_zero, + map_add' := λ _ _, rfl } + +instance : module ℂ (cusp_form Γ k) := +function.injective.module ℂ coe_hom fun_like.coe_injective (λ _ _, rfl) + +instance : inhabited (cusp_form Γ k) := ⟨0⟩ + +@[priority 99] +instance [cusp_form_class F Γ k] : modular_form_class F Γ k := +{ coe := fun_like.coe, + coe_injective' := fun_like.coe_injective', + slash_action_eq := cusp_form_class.slash_action_eq, + holo:= cusp_form_class.holo, + bdd_at_infty := λ _ _, (cusp_form_class.zero_at_infty _ _).bounded_at_filter} + +end cusp_form diff --git a/src/number_theory/modular_forms/congruence_subgroups.lean b/src/number_theory/modular_forms/congruence_subgroups.lean index 5dfd42d0e3b01..3426c626336ec 100644 --- a/src/number_theory/modular_forms/congruence_subgroups.lean +++ b/src/number_theory/modular_forms/congruence_subgroups.lean @@ -10,7 +10,7 @@ import group_theory.group_action.conj_act /-! # Congruence subgroups -This defines congruence subgroups of `SL(2,ℤ)` such as `Γ(N)`, `Γ₀(N)` and `Γ₁(N)` for `N` a +This defines congruence subgroups of `SL(2, ℤ)` such as `Γ(N)`, `Γ₀(N)` and `Γ₁(N)` for `N` a natural number. It also contains basic results about congruence subgroups. @@ -35,7 +35,7 @@ lemma SL_reduction_mod_hom_val (N : ℕ) (γ : SL(2, ℤ)) : ∀ (i j : fin 2), ((SLMOD(N) γ) : (matrix (fin 2) (fin 2) (zmod N))) i j = (((↑ₘγ i j) : ℤ) : zmod N) := λ i j, rfl -/--The full level `N` congruence subgroup of `SL(2,ℤ)` of matrices that reduce to the identity +/--The full level `N` congruence subgroup of `SL(2, ℤ)` of matrices that reduce to the identity modulo `N`.-/ def Gamma (N : ℕ) : subgroup SL(2, ℤ) := (SLMOD(N)).ker @@ -79,7 +79,7 @@ begin simp [h] } end -/--The congruence subgroup of `SL(2,ℤ)` of matrices whose lower left-hand entry reduces to zero +/--The congruence subgroup of `SL(2, ℤ)` of matrices whose lower left-hand entry reduces to zero modulo `N`. -/ def Gamma0 (N : ℕ) : subgroup SL(2, ℤ) := { carrier := { g : SL(2, ℤ) | ((↑ₘg 1 0 : ℤ) : zmod N) = 0 }, @@ -147,7 +147,7 @@ begin exact ha.2.1,} end -/--The congruence subgroup `Gamma1` of `SL(2,ℤ)` consisting of matrices whose bottom +/--The congruence subgroup `Gamma1` of `SL(2, ℤ)` consisting of matrices whose bottom row is congruent to `(0,1)` modulo `N`. -/ def Gamma1 (N : ℕ) : subgroup SL(2, ℤ) := subgroup.map (((Gamma0 N).subtype).comp (Gamma1' N).subtype) ⊤ @@ -185,7 +185,7 @@ end section congruence_subgroup -/--A congruence subgroup is a subgroup of `SL(2,ℤ)` which contains some `Gamma N` for some +/--A congruence subgroup is a subgroup of `SL(2, ℤ)` which contains some `Gamma N` for some `(N : ℕ+)`. -/ def is_congruence_subgroup (Γ : subgroup SL(2, ℤ)) : Prop := ∃ (N : ℕ+), Gamma N ≤ Γ diff --git a/src/number_theory/modular_forms/slash_actions.lean b/src/number_theory/modular_forms/slash_actions.lean index 25b616b3bd5b0..e572d30bb678e 100644 --- a/src/number_theory/modular_forms/slash_actions.lean +++ b/src/number_theory/modular_forms/slash_actions.lean @@ -25,22 +25,20 @@ local notation `GL(` n `, ` R `)`⁺ := matrix.GL_pos (fin n) R local notation `SL(` n `, ` R `)` := matrix.special_linear_group (fin n) R /--A general version of the slash action of the space of modular forms.-/ -class slash_action (β G α γ : Type*) [group G] [has_zero α] - [has_one α] [has_smul γ α] [has_add α] := +class slash_action (β G α γ : Type*) [group G] [has_zero α] [has_smul γ α] [has_add α] := (map : β → G → α → α) -(mul_zero : ∀ (k : β) (g : G), map k g 0 = 0) -(one_mul : ∀ (k : β) (a : α) , map k 1 a = a) +(zero_slash : ∀ (k : β) (g : G), map k g 0 = 0) +(slash_one : ∀ (k : β) (a : α) , map k 1 a = a) (right_action : ∀ (k : β) (g h : G) (a : α), map k h (map k g a) = map k (g * h) a ) (smul_action : ∀ (k : β) (g : G) (a : α) (z : γ), map k g (z • a) = z • (map k g a)) (add_action : ∀ (k : β) (g : G) (a b : α), map k g (a + b) = map k g a + map k g b) /--Slash_action induced by a monoid homomorphism.-/ -def monoid_hom_slash_action {β G H α γ : Type*} [group G] [has_zero α] - [has_one α] [has_smul γ α] [has_add α] [group H] [slash_action β G α γ] - (h : H →* G) : slash_action β H α γ := +def monoid_hom_slash_action {β G H α γ : Type*} [group G] [has_zero α] [has_smul γ α] [has_add α] + [group H] [slash_action β G α γ] (h : H →* G) : slash_action β H α γ := { map := λ k g, slash_action.map γ k (h g), - mul_zero := λ k g, slash_action.mul_zero k (h g), - one_mul := λ k a, by simp only [map_one, slash_action.one_mul], + zero_slash := λ k g, slash_action.zero_slash k (h g), + slash_one := λ k a, by simp only [map_one, slash_action.slash_one], right_action := λ k g gg a, by simp only [map_mul, slash_action.right_action], smul_action := λ _ _, slash_action.smul_action _ _, add_action := λ _ g _ _, slash_action.add_action _ (h g) _ _,} @@ -77,7 +75,7 @@ begin simp_rw [this, ← mul_assoc, ←mul_zpow], end -lemma slash_add (k : ℤ) (A : GL(2, ℝ)⁺) (f g : ℍ → ℂ) : +@[simp] lemma slash_add (k : ℤ) (A : GL(2, ℝ)⁺) (f g : ℍ → ℂ) : (f + g) ∣[k] A = (f ∣[k] A) + (g ∣[k] A) := begin ext1, @@ -85,11 +83,15 @@ begin ring, end -lemma slash_one (k : ℤ) (f : ℍ → ℂ) : (f ∣[k] 1) = f := +@[simp] lemma slash_one (k : ℤ) (f : ℍ → ℂ) : (f ∣[k] 1) = f := funext $ by simp [slash] -lemma smul_slash (k : ℤ) (A : GL(2, ℝ)⁺) (f : ℍ → ℂ) (c : ℂ) : (c • f) ∣[k] A = c • (f ∣[k] A) := +variables {α : Type*} [has_smul α ℂ] [is_scalar_tower α ℂ ℂ] + +@[simp] lemma smul_slash (k : ℤ) (A : GL(2, ℝ)⁺) (f : ℍ → ℂ) (c : α) : + (c • f) ∣[k] A = c • (f ∣[k] A) := begin + simp_rw [←smul_one_smul ℂ c f, ←smul_one_smul ℂ c (f ∣[k] A)], ext1, simp_rw slash, simp only [slash, algebra.id.smul_eq_mul, matrix.general_linear_group.coe_det_apply, @@ -97,13 +99,16 @@ begin ring, end -lemma neg_slash (k : ℤ) (A : GL(2, ℝ)⁺) (f : ℍ → ℂ) : (-f) ∣[k] A = - (f ∣[k] A) := +@[simp] lemma neg_slash (k : ℤ) (A : GL(2, ℝ)⁺) (f : ℍ → ℂ) : (-f) ∣[k] A = - (f ∣[k] A) := funext $ by simp [slash] +@[simp] lemma zero_slash (k : ℤ) (A : GL(2, ℝ)⁺) : (0 : ℍ → ℂ) ∣[k] A = 0 := +funext $ λ _, by simp only [slash, pi.zero_apply, zero_mul] + instance : slash_action ℤ GL(2, ℝ)⁺ (ℍ → ℂ) ℂ := { map := slash, - mul_zero := λ k g, funext $ λ _, by simp only [slash, pi.zero_apply, zero_mul], - one_mul := slash_one, + zero_slash := zero_slash, + slash_one := slash_one, right_action := slash_right_action, smul_action := smul_slash, add_action := slash_add } @@ -124,8 +129,8 @@ monoid_hom_slash_action (monoid_hom.comp (matrix.special_linear_group.to_GL_pos) local notation f `∣[`:73 k:0, A `]` :72 := slash_action.map ℂ k A f -/-- The constant function 1 is invariant under any subgroup of `SL(2, ℤ)`. -/ -lemma is_invariant_one (A : SL(2, ℤ)) : (1 : ℍ → ℂ) ∣[(0 : ℤ), A] = (1 : ℍ → ℂ) := +/-- The constant function 1 is invariant under any element of `SL(2, ℤ)`. -/ +@[simp] lemma is_invariant_one (A : SL(2, ℤ)) : (1 : ℍ → ℂ) ∣[(0 : ℤ), A] = (1 : ℍ → ℂ) := begin have : (((↑ₘ(A : GL(2,ℝ)⁺)).det) : ℝ) = 1, { simp only [coe_coe, @@ -173,7 +178,7 @@ begin ring, end -lemma mul_slash_SL2 (k1 k2 : ℤ) (A : SL(2, ℤ)) (f g : ℍ → ℂ) : +@[simp] lemma mul_slash_SL2 (k1 k2 : ℤ) (A : SL(2, ℤ)) (f g : ℍ → ℂ) : (f * g) ∣[k1 + k2, A] = (f ∣[k1, A]) * (g ∣[k2, A]) := calc (f * g) ∣[k1 + k2, (A : GL(2, ℝ)⁺)] = _ • (f ∣[k1, A]) * (g ∣[k2, A]) : mul_slash _ _ _ _ _ ... = (1:ℝ) • (f ∣[k1, A]) * (g ∣[k2, A]) : by simp [-matrix.special_linear_group.coe_matrix_coe] diff --git a/src/number_theory/modular_forms/slash_invariant_forms.lean b/src/number_theory/modular_forms/slash_invariant_forms.lean index d8218c5060efa..b9f183db102fc 100644 --- a/src/number_theory/modular_forms/slash_invariant_forms.lean +++ b/src/number_theory/modular_forms/slash_invariant_forms.lean @@ -109,7 +109,7 @@ instance has_add : has_add (slash_invariant_form Γ k) := instance has_zero : has_zero (slash_invariant_form Γ k) := ⟨ { to_fun := 0, - slash_action_eq' := slash_action.mul_zero _} ⟩ + slash_action_eq' := slash_action.zero_slash _} ⟩ @[simp] lemma coe_zero : ⇑(0 : slash_invariant_form Γ k) = (0 : ℍ → ℂ) := rfl @@ -161,6 +161,8 @@ instance : has_one (slash_invariant_form Γ 0) := ⟨ { to_fun := 1, slash_action_eq' := λ A, modular_form.is_invariant_one A } ⟩ +@[simp] lemma one_coe_eq_one : ((1 : slash_invariant_form Γ 0) : ℍ → ℂ) = 1 := rfl + instance : inhabited (slash_invariant_form Γ k) := ⟨0⟩ end slash_invariant_form diff --git a/src/order/filter/zero_and_bounded_at_filter.lean b/src/order/filter/zero_and_bounded_at_filter.lean index 3f51234f2975d..49cdbe93dfba0 100644 --- a/src/order/filter/zero_and_bounded_at_filter.lean +++ b/src/order/filter/zero_and_bounded_at_filter.lean @@ -28,53 +28,99 @@ open_locale topological_space def zero_at_filter [has_zero β] [topological_space β] (l : filter α) (f : α → β) : Prop := filter.tendsto f l (𝓝 0) -lemma zero_is_zero_at_filter [has_zero β] [topological_space β] (l : filter α) : zero_at_filter l - (0 : α → β) := tendsto_const_nhds +lemma zero_zero_at_filter [has_zero β] [topological_space β] (l : filter α) : + zero_at_filter l (0 : α → β) := +tendsto_const_nhds + +lemma zero_at_filter.add [topological_space β] [add_zero_class β] [has_continuous_add β] + {l : filter α} {f g : α → β} (hf : zero_at_filter l f) (hg : zero_at_filter l g) : + zero_at_filter l (f + g) := +by simpa using hf.add hg + +lemma zero_at_filter.neg [topological_space β] [add_group β] [has_continuous_neg β] {l : filter α} + {f : α → β} (hf : zero_at_filter l f) : + zero_at_filter l (-f) := +by simpa using hf.neg + +lemma zero_at_filter.smul {𝕜 : Type*} [topological_space 𝕜] [topological_space β] [has_zero 𝕜] + [has_zero β] [smul_with_zero 𝕜 β] [has_continuous_smul 𝕜 β] + {l : filter α} {f : α → β} (c : 𝕜) (hf : zero_at_filter l f) : + zero_at_filter l (c • f) := +by simpa using hf.const_smul c /-- `zero_at_filter_submodule l` is the submodule of `f : α → β` which tend to zero along `l`. -/ def zero_at_filter_submodule [topological_space β] [semiring β] [has_continuous_add β] [has_continuous_mul β] (l : filter α) : submodule β (α → β) := { carrier := zero_at_filter l, - zero_mem' := zero_is_zero_at_filter l, - add_mem' := by { intros a b ha hb, simpa using ha.add hb, }, - smul_mem' := by { intros c f hf, simpa using hf.const_mul c }, } + zero_mem' := zero_zero_at_filter l, + add_mem' := λ a b ha hb, ha.add hb, + smul_mem' := λ c f hf, hf.smul c } /-- `zero_at_filter_add_submonoid l` is the additive submonoid of `f : α → β` which tend to zero along `l`. -/ def zero_at_filter_add_submonoid [topological_space β] [add_zero_class β] [has_continuous_add β] (l : filter α) : add_submonoid (α → β) := { carrier := zero_at_filter l, - add_mem' := by { intros a b ha hb, simpa using ha.add hb }, - zero_mem' := zero_is_zero_at_filter l, } + add_mem' := λ a b ha hb, ha.add hb, + zero_mem' := zero_zero_at_filter l, } /-- If `l` is a filter on `α`, then a function `f: α → β` is `bounded_at_filter l` if `f =O[l] 1`. -/ -def bounded_at_filter [has_norm β] [has_one (α → β)] (l : filter α) (f : α → β) : Prop := -asymptotics.is_O l f (1 : α → β) +def bounded_at_filter [has_norm β] (l : filter α) (f : α → β) : Prop := +asymptotics.is_O l f (1 : α → ℝ) -lemma zero_at_filter.bounded_at_filter [normed_field β] {l : filter α} {f : α → β} +lemma zero_at_filter.bounded_at_filter [normed_add_comm_group β] {l : filter α} {f : α → β} (hf : zero_at_filter l f) : bounded_at_filter l f := -asymptotics.is_O_of_div_tendsto_nhds (by simp) _ (by { convert hf, ext1, simp, }) +begin + rw [zero_at_filter, ← asymptotics.is_o_const_iff (one_ne_zero' ℝ)] at hf, + exact hf.is_O, +end + +lemma const_bounded_at_filter [normed_field β] (l : filter α) (c : β) : + bounded_at_filter l (function.const α c : α → β) := +asymptotics.is_O_const_const c one_ne_zero l + +lemma bounded_at_filter.add [normed_add_comm_group β] {l : filter α} {f g : α → β} + (hf : bounded_at_filter l f) (hg : bounded_at_filter l g) : + bounded_at_filter l (f + g) := +by simpa using hf.add hg + +lemma bounded_at_filter.neg [normed_add_comm_group β] {l : filter α} {f : α → β} + (hf : bounded_at_filter l f) : + bounded_at_filter l (-f) := +hf.neg_left -lemma zero_is_bounded_at_filter [normed_field β] (l : filter α) : - bounded_at_filter l (0 : α → β) := -(zero_is_zero_at_filter l).bounded_at_filter +lemma bounded_at_filter.smul {𝕜 : Type*} [normed_field 𝕜] [normed_add_comm_group β] + [normed_space 𝕜 β] {l : filter α} {f : α → β} (c : 𝕜) (hf : bounded_at_filter l f) : + bounded_at_filter l (c • f) := +hf.const_smul_left c + +lemma bounded_at_filter.mul [normed_field β] {l : filter α} {f g : α → β} + (hf : bounded_at_filter l f) (hg : bounded_at_filter l g) : + bounded_at_filter l (f * g) := +begin + refine (hf.mul hg).trans _, + convert asymptotics.is_O_refl _ l, + ext x, + simp, +end /-- The submodule of functions that are bounded along a filter `l`. -/ def bounded_filter_submodule [normed_field β] (l : filter α) : submodule β (α → β) := { carrier := bounded_at_filter l, - zero_mem' := zero_is_bounded_at_filter l, - add_mem' := by { intros f g hf hg, simpa using hf.add hg, }, - smul_mem' := by { intros c f hf, simpa using hf.const_mul_left c }, } + zero_mem' := const_bounded_at_filter l 0, + add_mem' := λ f g hf hg, hf.add hg, + smul_mem' := λ c f hf, hf.smul c } /-- The subalgebra of functions that are bounded along a filter `l`. -/ def bounded_filter_subalgebra [normed_field β] (l : filter α) : subalgebra β (α → β) := begin refine submodule.to_subalgebra (bounded_filter_submodule l) _ (λ f g hf hg, _), - { simpa using asymptotics.is_O_const_mul_self (1 : β) (1 : α → β) l }, + { exact const_bounded_at_filter l (1:β) }, { simpa only [pi.one_apply, mul_one, norm_mul] using hf.mul hg, }, + end end filter From fc2ed6f838ce7c9b7c7171e58d78eaf7b438fb0e Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 12 Dec 2022 17:59:51 +0000 Subject: [PATCH 0020/1291] feat(logic/denumerable): Any two infinite countable types are equivalent (#17903) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit If `α` and `β` are both infinite and countable, then `nonempty (α ≃ β)`. I prove it using the `denumerable` API, but I also provide a lemma to prove it from the `cardinal` one. --- src/data/fintype/card.lean | 1 + src/logic/denumerable.lean | 16 +++++++++++++++- src/logic/encodable/basic.lean | 1 + src/set_theory/cardinal/basic.lean | 3 +++ 4 files changed, 20 insertions(+), 1 deletion(-) diff --git a/src/data/fintype/card.lean b/src/data/fintype/card.lean index 66f2c9edb8a8a..5f8f084613e0e 100644 --- a/src/data/fintype/card.lean +++ b/src/data/fintype/card.lean @@ -338,6 +338,7 @@ lemma finite_iff_nonempty_fintype (α : Type*) : ⟨λ h, let ⟨k, ⟨e⟩⟩ := @finite.exists_equiv_fin α h in ⟨fintype.of_equiv _ e.symm⟩, λ ⟨_⟩, by exactI infer_instance⟩ +/-- See also `nonempty_encodable`, `nonempty_denumerable`. -/ lemma nonempty_fintype (α : Type*) [finite α] : nonempty (fintype α) := (finite_iff_nonempty_fintype α).mp ‹_› diff --git a/src/logic/denumerable.lean b/src/logic/denumerable.lean index 7d5378d402e0b..f3510d58ae41f 100644 --- a/src/logic/denumerable.lean +++ b/src/logic/denumerable.lean @@ -21,6 +21,8 @@ This property already has a name, namely `α ≃ ℕ`, but here we are intereste typeclass. -/ +variables {α β : Type*} + /-- A denumerable type is (constructively) bijective with `ℕ`. Typeclass equivalent of `α ≃ ℕ`. -/ class denumerable (α : Type*) extends encodable α := (decode_inv : ∀ n, ∃ a ∈ decode n, encode a = n) @@ -30,7 +32,7 @@ open nat namespace denumerable section -variables {α : Type*} {β : Type*} [denumerable α] [denumerable β] +variables [denumerable α] [denumerable β] open encodable theorem decode_is_some (α) [denumerable α] (n : ℕ) : @@ -290,3 +292,15 @@ begin end end denumerable + +/-- See also `nonempty_encodable`, `nonempty_fintype`. -/ +lemma nonempty_denumerable (α : Type*) [countable α] [infinite α] : nonempty (denumerable α) := +(nonempty_encodable α).map $ λ h, by exactI denumerable.of_encodable_of_infinite _ + +instance nonempty_equiv_of_countable [countable α] [infinite α] [countable β] [infinite β] : + nonempty (α ≃ β) := +begin + casesI nonempty_denumerable α, + casesI nonempty_denumerable β, + exact ⟨(denumerable.eqv _).trans (denumerable.eqv _).symm⟩, +end diff --git a/src/logic/encodable/basic.lean b/src/logic/encodable/basic.lean index 1f30b2ecf509a..9d47356f8757d 100644 --- a/src/logic/encodable/basic.lean +++ b/src/logic/encodable/basic.lean @@ -335,6 +335,7 @@ nonempty.some $ let ⟨f, hf⟩ := exists_injective_nat α in ⟨of_inj f hf⟩ end encodable +/-- See also `nonempty_fintype`, `nonempty_denumerable`. -/ lemma nonempty_encodable (α : Type*) [countable α] : nonempty (encodable α) := ⟨encodable.of_countable _⟩ diff --git a/src/set_theory/cardinal/basic.lean b/src/set_theory/cardinal/basic.lean index cd08ffaf3e447..880b32b1fcaf4 100644 --- a/src/set_theory/cardinal/basic.lean +++ b/src/set_theory/cardinal/basic.lean @@ -1092,6 +1092,9 @@ by rw [← not_lt, lt_aleph_0_iff_finite, not_finite_iff_infinite] @[simp] lemma aleph_0_le_mk (α : Type u) [infinite α] : ℵ₀ ≤ #α := infinite_iff.1 ‹_› +@[simp] lemma mk_eq_aleph_0 (α : Type*) [countable α] [infinite α] : #α = ℵ₀ := +mk_le_aleph_0.antisymm $ aleph_0_le_mk _ + lemma denumerable_iff {α : Type u} : nonempty (denumerable α) ↔ #α = ℵ₀ := ⟨λ ⟨h⟩, mk_congr ((@denumerable.eqv α h).trans equiv.ulift.symm), λ h, by { cases quotient.exact h with f, exact ⟨denumerable.mk' $ f.trans equiv.ulift⟩ }⟩ From ac5a7cec422c3909db52e13dde2e729657d19b0e Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Mon, 12 Dec 2022 21:26:29 +0000 Subject: [PATCH 0021/1291] refactor(data/set/intervals/monotone): split (#17893) Also generalize&golf `monotone_on.exists_monotone_extension` and rename `order.monotone` to `order.monotone.basic`. Lean 4 version is leanprover-community/mathlib4#947 --- scripts/modules_used.lean | 2 +- src/algebra/covariant_and_contravariant.lean | 2 +- src/analysis/calculus/monotone.lean | 6 +- .../trigonometric/deriv.lean | 1 + src/data/int/basic.lean | 2 +- src/data/set/intervals/iso_Ioo.lean | 37 +++ src/data/set/intervals/monotone.lean | 222 +----------------- src/order/iterate.lean | 2 +- src/order/lattice.lean | 2 +- .../{monotone.lean => monotone/basic.lean} | 0 src/order/monotone/extension.lean | 55 +++++ src/order/monotone/odd.lean | 52 ++++ src/order/monotone/union.lean | 114 +++++++++ src/topology/tietze_extension.lean | 2 +- 14 files changed, 271 insertions(+), 228 deletions(-) create mode 100644 src/data/set/intervals/iso_Ioo.lean rename src/order/{monotone.lean => monotone/basic.lean} (100%) create mode 100644 src/order/monotone/extension.lean create mode 100644 src/order/monotone/odd.lean create mode 100644 src/order/monotone/union.lean diff --git a/scripts/modules_used.lean b/scripts/modules_used.lean index 8a7634f0fd805..43348c2a6b221 100644 --- a/scripts/modules_used.lean +++ b/scripts/modules_used.lean @@ -14,7 +14,7 @@ returns ``` order.synonym order.rel_classes -order.monotone +order.monotone.basic order.lattice order.heyting.basic order.bounded_order diff --git a/src/algebra/covariant_and_contravariant.lean b/src/algebra/covariant_and_contravariant.lean index a7918f1926195..3825be43ab134 100644 --- a/src/algebra/covariant_and_contravariant.lean +++ b/src/algebra/covariant_and_contravariant.lean @@ -6,7 +6,7 @@ Authors: Damiano Testa import algebra.group.defs import order.basic -import order.monotone +import order.monotone.basic /-! diff --git a/src/analysis/calculus/monotone.lean b/src/analysis/calculus/monotone.lean index d7f179ff05491..1bc2b999d16ad 100644 --- a/src/analysis/calculus/monotone.lean +++ b/src/analysis/calculus/monotone.lean @@ -6,6 +6,7 @@ Authors: Sébastien Gouëzel import measure_theory.measure.lebesgue import analysis.calculus.deriv import measure_theory.covering.one_dim +import order.monotone.extension /-! # Differentiability of monotone functions @@ -232,8 +233,9 @@ begin apply ae_of_mem_of_ae_of_mem_inter_Ioo, assume a b as bs hab, obtain ⟨g, hg, gf⟩ : ∃ (g : ℝ → ℝ), monotone g ∧ eq_on f g (s ∩ Icc a b) := - monotone_on.exists_monotone_extension (hf.mono (inter_subset_left s (Icc a b))) - ⟨⟨as, ⟨le_rfl, hab.le⟩⟩, λ x hx, hx.2.1⟩ ⟨⟨bs, ⟨hab.le, le_rfl⟩⟩, λ x hx, hx.2.2⟩, + (hf.mono (inter_subset_left s (Icc a b))).exists_monotone_extension + (hf.map_bdd_below (inter_subset_left _ _) ⟨a, λ x hx, hx.2.1, as⟩) + (hf.map_bdd_above (inter_subset_left _ _) ⟨b, λ x hx, hx.2.2, bs⟩), filter_upwards [hg.ae_differentiable_at] with x hx, assume h'x, apply hx.differentiable_within_at.congr_of_eventually_eq _ (gf ⟨h'x.1, h'x.2.1.le, h'x.2.2.le⟩), diff --git a/src/analysis/special_functions/trigonometric/deriv.lean b/src/analysis/special_functions/trigonometric/deriv.lean index 72f2967a2c419..30bf58d909902 100644 --- a/src/analysis/special_functions/trigonometric/deriv.lean +++ b/src/analysis/special_functions/trigonometric/deriv.lean @@ -3,6 +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, Benjamin Davidson -/ +import order.monotone.odd import analysis.special_functions.exp_deriv import analysis.special_functions.trigonometric.basic import data.set.intervals.monotone diff --git a/src/data/int/basic.lean b/src/data/int/basic.lean index 1fc1d36c9d984..36c01318130ec 100644 --- a/src/data/int/basic.lean +++ b/src/data/int/basic.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad -/ import data.nat.basic -import order.monotone +import order.monotone.basic /-! # Basic instances on the integers diff --git a/src/data/set/intervals/iso_Ioo.lean b/src/data/set/intervals/iso_Ioo.lean new file mode 100644 index 0000000000000..a3cfa15ccf236 --- /dev/null +++ b/src/data/set/intervals/iso_Ioo.lean @@ -0,0 +1,37 @@ +/- +Copyright (c) 2022 Yury Kudryashov. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yury Kudryashov +-/ +import order.monotone.odd +import tactic.field_simp + +/-! +# Order isomorphism between a linear ordered field and `(-1, 1)` + +In this file we provide an order isomorphism `order_iso_Ioo_neg_one_one` between the open interval +`(-1, 1)` in a linear ordered field and the whole field. +-/ + +open set + +/-- In a linear ordered field, the whole field is order isomorphic to the open interval `(-1, 1)`. +We consider the actual implementation to be a "black box", so it is irreducible. +-/ +@[irreducible] def order_iso_Ioo_neg_one_one (k : Type*) [linear_ordered_field k] : + k ≃o Ioo (-1 : k) 1 := +begin + refine strict_mono.order_iso_of_right_inverse _ _ (λ x, x / (1 - |x|)) _, + { refine cod_restrict (λ x, x / (1 + |x|)) _ (λ x, abs_lt.1 _), + have H : 0 < 1 + |x|, from (abs_nonneg x).trans_lt (lt_one_add _), + calc |x / (1 + |x|)| = |x| / (1 + |x|) : by rw [abs_div, abs_of_pos H] + ... < 1 : (div_lt_one H).2 (lt_one_add _) }, + { refine (strict_mono_of_odd_strict_mono_on_nonneg _ _).cod_restrict _, + { intro x, simp only [abs_neg, neg_div] }, + { rintros x (hx : 0 ≤ x) y (hy : 0 ≤ y) hxy, + simp [abs_of_nonneg, mul_add, mul_comm x y, div_lt_div_iff, + hx.trans_lt (lt_one_add _), hy.trans_lt (lt_one_add _), *] } }, + { refine λ x, subtype.ext _, + have : 0 < 1 - |(x : k)|, from sub_pos.2 (abs_lt.2 x.2), + field_simp [abs_div, this.ne', abs_of_pos this] } +end diff --git a/src/data/set/intervals/monotone.lean b/src/data/set/intervals/monotone.lean index d81696cf4d8a3..c9ce045e52930 100644 --- a/src/data/set/intervals/monotone.lean +++ b/src/data/set/intervals/monotone.lean @@ -4,235 +4,17 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import data.set.intervals.disjoint -import order.conditionally_complete_lattice.basic import order.succ_pred.basic -import algebra.order.field.basic -import tactic.field_simp /-! # Monotonicity on intervals -In this file we prove that a function is (strictly) monotone (or antitone) on a linear order `α` -provided that it is (strictly) monotone on `(-∞, a]` and on `[a, +∞)`. This is a special case -of a more general statement where one deduces monotonicity on a union from monotonicity on each -set. - -We deduce in `monotone_on.exists_monotone_extension` that a function which is monotone on a set -with a smallest and a largest element admits a monotone extension to the whole space. - -We also provide an order isomorphism `order_iso_Ioo_neg_one_one` between the open -interval `(-1, 1)` in a linear ordered field and the whole field. +In this file we prove that `set.Ici` etc are monotone/antitone functions. We also prove some lemmas +about functions monotone on intervals in `succ_order`s. -/ open set -section - -variables {α β : Type*} [linear_order α] [preorder β] {a : α} {f : α → β} - -/-- If `f` is strictly monotone both on `s` and `t`, with `s` to the left of `t` and the center -point belonging to both `s` and `t`, then `f` is strictly monotone on `s ∪ t` -/ -protected lemma strict_mono_on.union {s t : set α} {c : α} (h₁ : strict_mono_on f s) - (h₂ : strict_mono_on f t) (hs : is_greatest s c) (ht : is_least t c) : - strict_mono_on f (s ∪ t) := -begin - have A : ∀ x, x ∈ s ∪ t → x ≤ c → x ∈ s, - { assume x hx hxc, - cases hx, { exact hx }, - rcases eq_or_lt_of_le hxc with rfl|h'x, { exact hs.1 }, - exact (lt_irrefl _ (h'x.trans_le (ht.2 hx))).elim }, - have B : ∀ x, x ∈ s ∪ t → c ≤ x → x ∈ t, - { assume x hx hxc, - cases hx, swap, { exact hx }, - rcases eq_or_lt_of_le hxc with rfl|h'x, { exact ht.1 }, - exact (lt_irrefl _ (h'x.trans_le (hs.2 hx))).elim }, - assume x hx y hy hxy, - rcases lt_or_le x c with hxc|hcx, - { have xs : x ∈ s, from A _ hx hxc.le, - rcases lt_or_le y c with hyc|hcy, - { exact h₁ xs (A _ hy hyc.le) hxy }, - { exact (h₁ xs hs.1 hxc).trans_le (h₂.monotone_on ht.1 (B _ hy hcy) hcy) } }, - { have xt : x ∈ t, from B _ hx hcx, - have yt : y ∈ t, from B _ hy (hcx.trans hxy.le), - exact h₂ xt yt hxy } -end - -/-- If `f` is strictly monotone both on `(-∞, a]` and `[a, ∞)`, then it is strictly monotone on the -whole line. -/ -protected lemma strict_mono_on.Iic_union_Ici (h₁ : strict_mono_on f (Iic a)) - (h₂ : strict_mono_on f (Ici a)) : strict_mono f := -begin - rw [← strict_mono_on_univ, ← @Iic_union_Ici _ _ a], - exact strict_mono_on.union h₁ h₂ is_greatest_Iic is_least_Ici, -end - -/-- If `f` is strictly antitone both on `s` and `t`, with `s` to the left of `t` and the center -point belonging to both `s` and `t`, then `f` is strictly antitone on `s ∪ t` -/ -protected lemma strict_anti_on.union {s t : set α} {c : α} (h₁ : strict_anti_on f s) - (h₂ : strict_anti_on f t) (hs : is_greatest s c) (ht : is_least t c) : - strict_anti_on f (s ∪ t) := -(h₁.dual_right.union h₂.dual_right hs ht).dual_right - -/-- If `f` is strictly antitone both on `(-∞, a]` and `[a, ∞)`, then it is strictly antitone on the -whole line. -/ -protected lemma strict_anti_on.Iic_union_Ici (h₁ : strict_anti_on f (Iic a)) - (h₂ : strict_anti_on f (Ici a)) : strict_anti f := -(h₁.dual_right.Iic_union_Ici h₂.dual_right).dual_right - -/-- If `f` is monotone both on `s` and `t`, with `s` to the left of `t` and the center -point belonging to both `s` and `t`, then `f` is monotone on `s ∪ t` -/ -protected lemma monotone_on.union_right {s t : set α} {c : α} (h₁ : monotone_on f s) - (h₂ : monotone_on f t) (hs : is_greatest s c) (ht : is_least t c) : - monotone_on f (s ∪ t) := -begin - have A : ∀ x, x ∈ s ∪ t → x ≤ c → x ∈ s, - { assume x hx hxc, - cases hx, { exact hx }, - rcases eq_or_lt_of_le hxc with rfl|h'x, { exact hs.1 }, - exact (lt_irrefl _ (h'x.trans_le (ht.2 hx))).elim }, - have B : ∀ x, x ∈ s ∪ t → c ≤ x → x ∈ t, - { assume x hx hxc, - cases hx, swap, { exact hx }, - rcases eq_or_lt_of_le hxc with rfl|h'x, { exact ht.1 }, - exact (lt_irrefl _ (h'x.trans_le (hs.2 hx))).elim }, - assume x hx y hy hxy, - rcases lt_or_le x c with hxc|hcx, - { have xs : x ∈ s, from A _ hx hxc.le, - rcases lt_or_le y c with hyc|hcy, - { exact h₁ xs (A _ hy hyc.le) hxy }, - { exact (h₁ xs hs.1 hxc.le).trans (h₂ ht.1 (B _ hy hcy) hcy) } }, - { have xt : x ∈ t, from B _ hx hcx, - have yt : y ∈ t, from B _ hy (hcx.trans hxy), - exact h₂ xt yt hxy } -end - -/-- If `f` is monotone both on `(-∞, a]` and `[a, ∞)`, then it is monotone on the whole line. -/ -protected lemma monotone_on.Iic_union_Ici (h₁ : monotone_on f (Iic a)) - (h₂ : monotone_on f (Ici a)) : monotone f := -begin - rw [← monotone_on_univ, ← @Iic_union_Ici _ _ a], - exact monotone_on.union_right h₁ h₂ is_greatest_Iic is_least_Ici -end - -/-- If `f` is antitone both on `s` and `t`, with `s` to the left of `t` and the center -point belonging to both `s` and `t`, then `f` is antitone on `s ∪ t` -/ -protected lemma antitone_on.union_right {s t : set α} {c : α} (h₁ : antitone_on f s) - (h₂ : antitone_on f t) (hs : is_greatest s c) (ht : is_least t c) : - antitone_on f (s ∪ t) := -(h₁.dual_right.union_right h₂.dual_right hs ht).dual_right - -/-- If `f` is antitone both on `(-∞, a]` and `[a, ∞)`, then it is antitone on the whole line. -/ -protected lemma antitone_on.Iic_union_Ici (h₁ : antitone_on f (Iic a)) - (h₂ : antitone_on f (Ici a)) : antitone f := -(h₁.dual_right.Iic_union_Ici h₂.dual_right).dual_right - -/-- If a function is monotone on a set `s`, then it admits a monotone extension to the whole space -provided `s` has a least element `a` and a greatest element `b`. -/ -lemma monotone_on.exists_monotone_extension {β : Type*} [conditionally_complete_linear_order β] - {f : α → β} {s : set α} (h : monotone_on f s) {a b : α} - (ha : is_least s a) (hb : is_greatest s b) : - ∃ g : α → β, monotone g ∧ eq_on f g s := -begin - /- The extension is defined by `f x = f a` for `x ≤ a`, and `f x` is the supremum of the values - of `f` to the left of `x` for `x ≥ a`. -/ - have aleb : a ≤ b := hb.2 ha.1, - have H : ∀ x ∈ s, f x = Sup (f '' (Icc a x ∩ s)), - { assume x xs, - have xmem : x ∈ Icc a x ∩ s := ⟨⟨ha.2 xs, le_rfl⟩, xs⟩, - have H : ∀ z, z ∈ f '' (Icc a x ∩ s) → z ≤ f x, - { rintros _ ⟨z, ⟨⟨az, zx⟩, zs⟩, rfl⟩, - exact h zs xs zx }, - apply le_antisymm, - { exact le_cSup ⟨f x, H⟩ (mem_image_of_mem _ xmem) }, - { exact cSup_le (nonempty_image_iff.2 ⟨x, xmem⟩) H } }, - let g := λ x, if x ≤ a then f a else Sup (f '' (Icc a x ∩ s)), - have hfg : eq_on f g s, - { assume x xs, - dsimp only [g], - by_cases hxa : x ≤ a, - { have : x = a, from le_antisymm hxa (ha.2 xs), - simp only [if_true, this, le_refl] }, - rw [if_neg hxa], - exact H x xs }, - have M1 : monotone_on g (Iic a), - { rintros x (hx : x ≤ a) y (hy : y ≤ a) hxy, - dsimp only [g], - simp only [hx, hy, if_true] }, - have g_eq : ∀ x ∈ Ici a, g x = Sup (f '' (Icc a x ∩ s)), - { rintros x ax, - dsimp only [g], - by_cases hxa : x ≤ a, - { have : x = a := le_antisymm hxa ax, - simp_rw [hxa, if_true, H a ha.1, this] }, - simp only [hxa, if_false], }, - have M2 : monotone_on g (Ici a), - { rintros x ax y ay hxy, - rw [g_eq x ax, g_eq y ay], - apply cSup_le_cSup, - { refine ⟨f b, _⟩, - rintros _ ⟨z, ⟨⟨az, zy⟩, zs⟩, rfl⟩, - exact h zs hb.1 (hb.2 zs) }, - { exact ⟨f a, mem_image_of_mem _ ⟨⟨le_rfl, ax⟩, ha.1⟩⟩ }, - { apply image_subset, - apply inter_subset_inter_left, - exact Icc_subset_Icc le_rfl hxy } }, - exact ⟨g, M1.Iic_union_Ici M2, hfg⟩, -end - -/-- If a function is antitone on a set `s`, then it admits an antitone extension to the whole space -provided `s` has a least element `a` and a greatest element `b`. -/ -lemma antitone_on.exists_antitone_extension {β : Type*} [conditionally_complete_linear_order β] - {f : α → β} {s : set α} (h : antitone_on f s) {a b : α} - (ha : is_least s a) (hb : is_greatest s b) : - ∃ g : α → β, antitone g ∧ eq_on f g s := -h.dual_right.exists_monotone_extension ha hb - -end - -section ordered_group - -variables {G H : Type*} [linear_ordered_add_comm_group G] [ordered_add_comm_group H] - -lemma strict_mono_of_odd_strict_mono_on_nonneg {f : G → H} (h₁ : ∀ x, f (-x) = -f x) - (h₂ : strict_mono_on f (Ici 0)) : - strict_mono f := -begin - refine strict_mono_on.Iic_union_Ici (λ x hx y hy hxy, neg_lt_neg_iff.1 _) h₂, - rw [← h₁, ← h₁], - exact h₂ (neg_nonneg.2 hy) (neg_nonneg.2 hx) (neg_lt_neg hxy) -end - -lemma monotone_of_odd_of_monotone_on_nonneg {f : G → H} (h₁ : ∀ x, f (-x) = -f x) - (h₂ : monotone_on f (Ici 0)) : monotone f := -begin - refine monotone_on.Iic_union_Ici (λ x hx y hy hxy, neg_le_neg_iff.1 _) h₂, - rw [← h₁, ← h₁], - exact h₂ (neg_nonneg.2 hy) (neg_nonneg.2 hx) (neg_le_neg hxy) -end - -end ordered_group - -/-- In a linear ordered field, the whole field is order isomorphic to the open interval `(-1, 1)`. -We consider the actual implementation to be a "black box", so it is irreducible. --/ -@[irreducible] def order_iso_Ioo_neg_one_one (k : Type*) [linear_ordered_field k] : - k ≃o Ioo (-1 : k) 1 := -begin - refine strict_mono.order_iso_of_right_inverse _ _ (λ x, x / (1 - |x|)) _, - { refine cod_restrict (λ x, x / (1 + |x|)) _ (λ x, abs_lt.1 _), - have H : 0 < 1 + |x|, from (abs_nonneg x).trans_lt (lt_one_add _), - calc |x / (1 + |x|)| = |x| / (1 + |x|) : by rw [abs_div, abs_of_pos H] - ... < 1 : (div_lt_one H).2 (lt_one_add _) }, - { refine (strict_mono_of_odd_strict_mono_on_nonneg _ _).cod_restrict _, - { intro x, simp only [abs_neg, neg_div] }, - { rintros x (hx : 0 ≤ x) y (hy : 0 ≤ y) hxy, - simp [abs_of_nonneg, mul_add, mul_comm x y, div_lt_div_iff, - hx.trans_lt (lt_one_add _), hy.trans_lt (lt_one_add _), *] } }, - { refine λ x, subtype.ext _, - have : 0 < 1 - |(x : k)|, from sub_pos.2 (abs_lt.2 x.2), - field_simp [abs_div, this.ne', abs_of_pos this] } -end - section Ixx variables {α β : Type*} [preorder α] [preorder β] {f g : α → β} {s : set α} diff --git a/src/order/iterate.lean b/src/order/iterate.lean index 79346254361e3..70921d0907b83 100644 --- a/src/order/iterate.lean +++ b/src/order/iterate.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury G. Kudryashov -/ import logic.function.iterate -import order.monotone +import order.monotone.basic /-! # Inequalities on iterates diff --git a/src/order/lattice.lean b/src/order/lattice.lean index fd452479a780c..7961e75b4779d 100644 --- a/src/order/lattice.lean +++ b/src/order/lattice.lean @@ -3,7 +3,7 @@ Copyright (c) 2017 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl -/ -import order.monotone +import order.monotone.basic import tactic.simps import tactic.pi_instances diff --git a/src/order/monotone.lean b/src/order/monotone/basic.lean similarity index 100% rename from src/order/monotone.lean rename to src/order/monotone/basic.lean diff --git a/src/order/monotone/extension.lean b/src/order/monotone/extension.lean new file mode 100644 index 0000000000000..ae27ab1e97a6b --- /dev/null +++ b/src/order/monotone/extension.lean @@ -0,0 +1,55 @@ +/- +Copyright (c) 2022 Sébastien Gouëzel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sébastien Gouëzel, Yury Kudryashov +-/ +import order.conditionally_complete_lattice.basic + +/-! +# Extension of a monotone function from a set to the whole space + +In this file we prove that if a function is monotone and is bounded on a set `s`, then it admits a +monotone extension to the whole space. +-/ + +open set + +variables {α β : Type*} [linear_order α] [conditionally_complete_linear_order β] + {f : α → β} {s : set α} {a b : α} + +/-- If a function is monotone and is bounded on a set `s`, then it admits a monotone extension to +the whole space. -/ +lemma monotone_on.exists_monotone_extension (h : monotone_on f s) (hl : bdd_below (f '' s)) + (hu : bdd_above (f '' s)) : + ∃ g : α → β, monotone g ∧ eq_on f g s := +begin + /- The extension is defined by `f x = f a` for `x ≤ a`, and `f x` is the supremum of the values + of `f` to the left of `x` for `x ≥ a`. -/ + classical, + rcases hl with ⟨a, ha⟩, + have hu' : ∀ x, bdd_above (f '' (Iic x ∩ s)), + from λ x, hu.mono (image_subset _ (inter_subset_right _ _)), + set g : α → β := λ x, if disjoint (Iic x) s then a else Sup (f '' (Iic x ∩ s)), + have hgs : eq_on f g s, + { intros x hx, + simp only [g], + have : is_greatest (Iic x ∩ s) x, from ⟨⟨right_mem_Iic, hx⟩, λ y hy, hy.1⟩, + rw [if_neg this.nonempty.not_disjoint, + ((h.mono $ inter_subset_right _ _).map_is_greatest this).cSup_eq] }, + refine ⟨g, λ x y hxy, _, hgs⟩, + by_cases hx : disjoint (Iic x) s; by_cases hy : disjoint (Iic y) s; + simp only [g, if_pos, if_neg, not_false_iff, *], + { rcases not_disjoint_iff_nonempty_inter.1 hy with ⟨z, hz⟩, + exact le_cSup_of_le (hu' _) (mem_image_of_mem _ hz) (ha $ mem_image_of_mem _ hz.2) }, + { exact (hx $ hy.mono_left $ Iic_subset_Iic.2 hxy).elim }, + { rw [not_disjoint_iff_nonempty_inter] at hx hy, + refine cSup_le_cSup (hu' _) (hx.image _) (image_subset _ _), + exact inter_subset_inter_left _ (Iic_subset_Iic.2 hxy) }, +end + +/-- If a function is antitone and is bounded on a set `s`, then it admits an antitone extension to +the whole space. -/ +lemma antitone_on.exists_antitone_extension (h : antitone_on f s) (hl : bdd_below (f '' s)) + (hu : bdd_above (f '' s)) : + ∃ g : α → β, antitone g ∧ eq_on f g s := +h.dual_right.exists_monotone_extension hu hl diff --git a/src/order/monotone/odd.lean b/src/order/monotone/odd.lean new file mode 100644 index 0000000000000..ca04a30e0a40e --- /dev/null +++ b/src/order/monotone/odd.lean @@ -0,0 +1,52 @@ +/- +Copyright (c) 2022 Yury Kudryashov. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yury Kudryashov +-/ +import order.monotone.union +import algebra.order.group.instances + +/-! +# Monotonicity of odd functions + +An odd function on a linear ordered additive commutative group `G` is monotone on the whole group +provided that is is monotone on `set.Ici 0`, see `monotone_of_odd_of_monotone_on_nonneg`. We also +prove versions of this lemma for `antitone`, `strict_mono`, and `strict_anti`. +-/ + +open set +variables {G H : Type*} [linear_ordered_add_comm_group G] [ordered_add_comm_group H] + +/-- An odd function on a linear ordered additive commutative group is strictly monotone on the whole +group provided that it is strictly monotone on `set.Ici 0`. -/ +lemma strict_mono_of_odd_strict_mono_on_nonneg {f : G → H} (h₁ : ∀ x, f (-x) = -f x) + (h₂ : strict_mono_on f (Ici 0)) : + strict_mono f := +begin + refine strict_mono_on.Iic_union_Ici (λ x hx y hy hxy, neg_lt_neg_iff.1 _) h₂, + rw [← h₁, ← h₁], + exact h₂ (neg_nonneg.2 hy) (neg_nonneg.2 hx) (neg_lt_neg hxy) +end + +/-- An odd function on a linear ordered additive commutative group is strictly antitone on the whole +group provided that it is strictly antitone on `set.Ici 0`. -/ +lemma strict_anti_of_odd_strict_anti_on_nonneg {f : G → H} (h₁ : ∀ x, f (-x) = -f x) + (h₂ : strict_anti_on f (Ici 0)) : + strict_anti f := +@strict_mono_of_odd_strict_mono_on_nonneg G Hᵒᵈ _ _ _ h₁ h₂ + +/-- An odd function on a linear ordered additive commutative group is monotone on the whole group +provided that it is monotone on `set.Ici 0`. -/ +lemma monotone_of_odd_of_monotone_on_nonneg {f : G → H} (h₁ : ∀ x, f (-x) = -f x) + (h₂ : monotone_on f (Ici 0)) : monotone f := +begin + refine monotone_on.Iic_union_Ici (λ x hx y hy hxy, neg_le_neg_iff.1 _) h₂, + rw [← h₁, ← h₁], + exact h₂ (neg_nonneg.2 hy) (neg_nonneg.2 hx) (neg_le_neg hxy) +end + +/-- An odd function on a linear ordered additive commutative group is antitone on the whole group +provided that it is monotone on `set.Ici 0`. -/ +lemma antitone_of_odd_of_monotone_on_nonneg {f : G → H} (h₁ : ∀ x, f (-x) = -f x) + (h₂ : antitone_on f (Ici 0)) : antitone f := +@monotone_of_odd_of_monotone_on_nonneg G Hᵒᵈ _ _ _ h₁ h₂ diff --git a/src/order/monotone/union.lean b/src/order/monotone/union.lean new file mode 100644 index 0000000000000..56bb83e27f15e --- /dev/null +++ b/src/order/monotone/union.lean @@ -0,0 +1,114 @@ +/- +Copyright (c) 2022 Yury Kudryashov. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yury Kudryashov, Sébastien Gouëzel +-/ +import order.bounds.basic + +/-! +# Monotonicity on intervals + +In this file we prove that a function is (strictly) monotone (or antitone) on a linear order `α` +provided that it is (strictly) monotone on `(-∞, a]` and on `[a, +∞)`. This is a special case +of a more general statement where one deduces monotonicity on a union from monotonicity on each +set. +-/ + +open set +variables {α β : Type*} [linear_order α] [preorder β] {a : α} {f : α → β} + +/-- If `f` is strictly monotone both on `s` and `t`, with `s` to the left of `t` and the center +point belonging to both `s` and `t`, then `f` is strictly monotone on `s ∪ t` -/ +protected lemma strict_mono_on.union {s t : set α} {c : α} (h₁ : strict_mono_on f s) + (h₂ : strict_mono_on f t) (hs : is_greatest s c) (ht : is_least t c) : + strict_mono_on f (s ∪ t) := +begin + have A : ∀ x, x ∈ s ∪ t → x ≤ c → x ∈ s, + { assume x hx hxc, + cases hx, { exact hx }, + rcases eq_or_lt_of_le hxc with rfl|h'x, { exact hs.1 }, + exact (lt_irrefl _ (h'x.trans_le (ht.2 hx))).elim }, + have B : ∀ x, x ∈ s ∪ t → c ≤ x → x ∈ t, + { assume x hx hxc, + cases hx, swap, { exact hx }, + rcases eq_or_lt_of_le hxc with rfl|h'x, { exact ht.1 }, + exact (lt_irrefl _ (h'x.trans_le (hs.2 hx))).elim }, + assume x hx y hy hxy, + rcases lt_or_le x c with hxc|hcx, + { have xs : x ∈ s, from A _ hx hxc.le, + rcases lt_or_le y c with hyc|hcy, + { exact h₁ xs (A _ hy hyc.le) hxy }, + { exact (h₁ xs hs.1 hxc).trans_le (h₂.monotone_on ht.1 (B _ hy hcy) hcy) } }, + { have xt : x ∈ t, from B _ hx hcx, + have yt : y ∈ t, from B _ hy (hcx.trans hxy.le), + exact h₂ xt yt hxy } +end + +/-- If `f` is strictly monotone both on `(-∞, a]` and `[a, ∞)`, then it is strictly monotone on the +whole line. -/ +protected lemma strict_mono_on.Iic_union_Ici (h₁ : strict_mono_on f (Iic a)) + (h₂ : strict_mono_on f (Ici a)) : strict_mono f := +begin + rw [← strict_mono_on_univ, ← @Iic_union_Ici _ _ a], + exact strict_mono_on.union h₁ h₂ is_greatest_Iic is_least_Ici, +end + +/-- If `f` is strictly antitone both on `s` and `t`, with `s` to the left of `t` and the center +point belonging to both `s` and `t`, then `f` is strictly antitone on `s ∪ t` -/ +protected lemma strict_anti_on.union {s t : set α} {c : α} (h₁ : strict_anti_on f s) + (h₂ : strict_anti_on f t) (hs : is_greatest s c) (ht : is_least t c) : + strict_anti_on f (s ∪ t) := +(h₁.dual_right.union h₂.dual_right hs ht).dual_right + +/-- If `f` is strictly antitone both on `(-∞, a]` and `[a, ∞)`, then it is strictly antitone on the +whole line. -/ +protected lemma strict_anti_on.Iic_union_Ici (h₁ : strict_anti_on f (Iic a)) + (h₂ : strict_anti_on f (Ici a)) : strict_anti f := +(h₁.dual_right.Iic_union_Ici h₂.dual_right).dual_right + +/-- If `f` is monotone both on `s` and `t`, with `s` to the left of `t` and the center +point belonging to both `s` and `t`, then `f` is monotone on `s ∪ t` -/ +protected lemma monotone_on.union_right {s t : set α} {c : α} (h₁ : monotone_on f s) + (h₂ : monotone_on f t) (hs : is_greatest s c) (ht : is_least t c) : + monotone_on f (s ∪ t) := +begin + have A : ∀ x, x ∈ s ∪ t → x ≤ c → x ∈ s, + { assume x hx hxc, + cases hx, { exact hx }, + rcases eq_or_lt_of_le hxc with rfl|h'x, { exact hs.1 }, + exact (lt_irrefl _ (h'x.trans_le (ht.2 hx))).elim }, + have B : ∀ x, x ∈ s ∪ t → c ≤ x → x ∈ t, + { assume x hx hxc, + cases hx, swap, { exact hx }, + rcases eq_or_lt_of_le hxc with rfl|h'x, { exact ht.1 }, + exact (lt_irrefl _ (h'x.trans_le (hs.2 hx))).elim }, + assume x hx y hy hxy, + rcases lt_or_le x c with hxc|hcx, + { have xs : x ∈ s, from A _ hx hxc.le, + rcases lt_or_le y c with hyc|hcy, + { exact h₁ xs (A _ hy hyc.le) hxy }, + { exact (h₁ xs hs.1 hxc.le).trans (h₂ ht.1 (B _ hy hcy) hcy) } }, + { have xt : x ∈ t, from B _ hx hcx, + have yt : y ∈ t, from B _ hy (hcx.trans hxy), + exact h₂ xt yt hxy } +end + +/-- If `f` is monotone both on `(-∞, a]` and `[a, ∞)`, then it is monotone on the whole line. -/ +protected lemma monotone_on.Iic_union_Ici (h₁ : monotone_on f (Iic a)) + (h₂ : monotone_on f (Ici a)) : monotone f := +begin + rw [← monotone_on_univ, ← @Iic_union_Ici _ _ a], + exact monotone_on.union_right h₁ h₂ is_greatest_Iic is_least_Ici +end + +/-- If `f` is antitone both on `s` and `t`, with `s` to the left of `t` and the center +point belonging to both `s` and `t`, then `f` is antitone on `s ∪ t` -/ +protected lemma antitone_on.union_right {s t : set α} {c : α} (h₁ : antitone_on f s) + (h₂ : antitone_on f t) (hs : is_greatest s c) (ht : is_least t c) : + antitone_on f (s ∪ t) := +(h₁.dual_right.union_right h₂.dual_right hs ht).dual_right + +/-- If `f` is antitone both on `(-∞, a]` and `[a, ∞)`, then it is antitone on the whole line. -/ +protected lemma antitone_on.Iic_union_Ici (h₁ : antitone_on f (Iic a)) + (h₂ : antitone_on f (Ici a)) : antitone f := +(h₁.dual_right.Iic_union_Ici h₂.dual_right).dual_right diff --git a/src/topology/tietze_extension.lean b/src/topology/tietze_extension.lean index 89bd87a13f8d9..def8be075ac0c 100644 --- a/src/topology/tietze_extension.lean +++ b/src/topology/tietze_extension.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Yury G. Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury G. Kudryashov -/ -import data.set.intervals.monotone +import data.set.intervals.iso_Ioo import topology.algebra.order.monotone_continuity import topology.urysohns_bounded From a5c51665cbd53e523ae656c089958304805973ed Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 12 Dec 2022 21:26:30 +0000 Subject: [PATCH 0022/1291] fix(ci): do not fail when detecting ported files more than 40 commits old (#17914) The previous action worked by using `git diff`, but this fails on shallow checkouts. This version hits the github API. --- .github/workflows/add_ported_warnings.yml | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/.github/workflows/add_ported_warnings.yml b/.github/workflows/add_ported_warnings.yml index b80932bc22d8a..c620e17518c81 100644 --- a/.github/workflows/add_ported_warnings.yml +++ b/.github/workflows/add_ported_warnings.yml @@ -22,8 +22,8 @@ jobs: # TODO: is this really faster than just calling git from python? - name: Get changed files id: changed-files - uses: tj-actions/changed-files@v34 + uses: jitterbit/get-changed-files@v1 - name: run the script run: | - python scripts/detect_ported_files.py ${{ steps.changed-files.outputs.all_changed_files }} + python scripts/detect_ported_files.py ${{ steps.changed-files.outputs.all }} From 2be0e5fd803921991ae793370d2e135b9469d65f Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Tue, 13 Dec 2022 00:40:32 +0000 Subject: [PATCH 0023/1291] feat(group_theory/subgroup/basic): add lemmas about `subgroup.map` (#17894) Add lemmas about comparison of `map f H` and `map f K` that don't assume `injective f`. --- src/group_theory/subgroup/basic.lean | 16 ++++++++++++++++ 1 file changed, 16 insertions(+) diff --git a/src/group_theory/subgroup/basic.lean b/src/group_theory/subgroup/basic.lean index 857cc9496029e..23362e9ec2150 100644 --- a/src/group_theory/subgroup/basic.lean +++ b/src/group_theory/subgroup/basic.lean @@ -2332,6 +2332,22 @@ lemma comap_map_eq_self_of_injective {f : G →* N} (h : function.injective f) ( comap f (map f H) = H := comap_map_eq_self (((ker_eq_bot_iff _).mpr h).symm ▸ bot_le) +@[to_additive] +lemma map_le_map_iff {f : G →* N} {H K : subgroup G} : H.map f ≤ K.map f ↔ H ≤ K ⊔ f.ker := +by rw [map_le_iff_le_comap, comap_map_eq] + +@[to_additive] lemma map_le_map_iff' {f : G →* N} {H K : subgroup G} : + H.map f ≤ K.map f ↔ H ⊔ f.ker ≤ K ⊔ f.ker := +by simp only [map_le_map_iff, sup_le_iff, le_sup_right, and_true] + +@[to_additive] lemma map_eq_map_iff {f : G →* N} {H K : subgroup G} : + H.map f = K.map f ↔ H ⊔ f.ker = K ⊔ f.ker := +by simp only [le_antisymm_iff, map_le_map_iff'] + +@[to_additive] lemma map_eq_range_iff {f : G →* N} {H : subgroup G} : + H.map f = f.range ↔ codisjoint H f.ker := +by rw [f.range_eq_map, map_eq_map_iff, codisjoint_iff, top_sup_eq] + @[to_additive] lemma map_le_map_iff_of_injective {f : G →* N} (hf : function.injective f) {H K : subgroup G} : H.map f ≤ K.map f ↔ H ≤ K := From b0cd2099301ff78659bfcd74cd917c64098e1a9f Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 13 Dec 2022 00:40:33 +0000 Subject: [PATCH 0024/1291] =?UTF-8?q?feat(dynamics/fixed=5Fpoints/basic):?= =?UTF-8?q?=20`x`=20is=20fixed=20by=20`f=E2=81=BF`=20if=20it=20is=20fixed?= =?UTF-8?q?=20by=20`f`=20(#17908)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Rephrase `function.is_fixed_pt.iterate` and `equiv.is_fixed_pt.symm` in new lemmas that use the multiplicative structure on `perm α`. Rename `equiv.is_fixed_pt.symm`/`equiv.is_fixed_pt.zpow` to `function.is_fixed_pt.equiv_symm`/`function.is_fixed_pt.perm_zpow` so that dot notation is possible. --- src/dynamics/fixed_points/basic.lean | 33 ++++++++++++-------------- src/group_theory/order_of_element.lean | 2 +- 2 files changed, 16 insertions(+), 19 deletions(-) diff --git a/src/dynamics/fixed_points/basic.lean b/src/dynamics/fixed_points/basic.lean index e25efac519929..8ab7ed997ebbb 100644 --- a/src/dynamics/fixed_points/basic.lean +++ b/src/dynamics/fixed_points/basic.lean @@ -22,9 +22,11 @@ We also prove some simple lemmas about `is_fixed_pt` and `∘`, `iterate`, and ` fixed point -/ +open equiv + universes u v -variables {α : Type u} {β : Type v} {f fa g : α → α} {x y : α} {fb : β → β} {m n k : ℕ} {e : α ≃ α} +variables {α : Type u} {β : Type v} {f fa g : α → α} {x y : α} {fb : β → β} {m n k : ℕ} {e : perm α} namespace function @@ -76,6 +78,18 @@ lemma preimage_iterate {s : set α} (h : is_fixed_pt (set.preimage f) s) (n : is_fixed_pt (set.preimage (f^[n])) s := by { rw set.preimage_iterate_eq, exact h.iterate n, } +protected lemma equiv_symm (h : is_fixed_pt e x) : is_fixed_pt e.symm x := +h.to_left_inverse e.left_inverse_symm + +protected lemma perm_inv (h : is_fixed_pt e x) : is_fixed_pt ⇑(e⁻¹) x := h.equiv_symm + +protected lemma perm_pow (h : is_fixed_pt e x) (n : ℕ) : is_fixed_pt ⇑(e ^ n) x := +by { rw ←equiv.perm.iterate_eq_pow, exact h.iterate _ } + +protected lemma perm_zpow (h : is_fixed_pt e x) : ∀ n : ℤ, is_fixed_pt ⇑(e ^ n) x +| (int.of_nat n) := h.perm_pow _ +| (int.neg_succ_of_nat n) := (h.perm_pow $ n + 1).perm_inv + end is_fixed_pt @[simp] lemma injective.is_fixed_pt_apply_iff (hf : injective f) {x : α} : @@ -143,20 +157,3 @@ lemma commute.right_bij_on_fixed_pts_comp (h : commute f g) : by simpa only [h.comp_eq] using bij_on_fixed_pts_comp f g end function - -namespace equiv.is_fixed_pt - -protected lemma symm (h : function.is_fixed_pt e x) : function.is_fixed_pt e.symm x := -h.to_left_inverse e.left_inverse_symm - -protected lemma zpow (h : function.is_fixed_pt e x) (n : ℤ) : function.is_fixed_pt ⇑(e^n) x := -begin - cases n, - { rw [int.of_nat_eq_coe, zpow_coe_nat, ← equiv.perm.iterate_eq_pow], - exact h.iterate n, }, - { change function.is_fixed_pt ⇑(e^(-(↑(n + 1) : ℤ))) x, - rw [zpow_neg, zpow_coe_nat, ← inv_pow, ← equiv.perm.iterate_eq_pow, equiv.perm.inv_def], - exact (equiv.is_fixed_pt.symm h).iterate (n + 1), }, -end - -end equiv.is_fixed_pt diff --git a/src/group_theory/order_of_element.lean b/src/group_theory/order_of_element.lean index eb667892a80e1..06b34f5afe778 100644 --- a/src/group_theory/order_of_element.lean +++ b/src/group_theory/order_of_element.lean @@ -449,7 +449,7 @@ begin obtain ⟨k, rfl⟩ := subgroup.mem_zpowers_iff.mp hx, rw [← mul_action.to_perm_apply, ← mul_action.to_perm_hom_apply, monoid_hom.map_zpow _ y k, mul_action.to_perm_hom_apply], - exact equiv.is_fixed_pt.zpow hs k, + exact function.is_fixed_pt.perm_zpow hs k, end lemma vadd_eq_self_of_mem_zmultiples {α G : Type*} [add_group G] [add_action G α] {x y : G} From d8615108625fe6dfb23d276fece6d1a5b989f230 Mon Sep 17 00:00:00 2001 From: Michail Karatarakis Date: Tue, 13 Dec 2022 03:30:54 +0000 Subject: [PATCH 0025/1291] feat(ring_theory/valuation/valuation_subring): Add `valuation_subring.comap` (#17310) Some definitions required for the definition of Frobenius elements. Co-authored-by: Michail Karatarakis <40603357+mkaratarakis@users.noreply.github.com> Co-authored-by: Eric Wieser --- src/ring_theory/subsemiring/basic.lean | 15 ++++++-- .../valuation/valuation_subring.lean | 35 +++++++++++++++++-- 2 files changed, 45 insertions(+), 5 deletions(-) diff --git a/src/ring_theory/subsemiring/basic.lean b/src/ring_theory/subsemiring/basic.lean index e55183cc3e227..3396929e837c8 100644 --- a/src/ring_theory/subsemiring/basic.lean +++ b/src/ring_theory/subsemiring/basic.lean @@ -881,9 +881,9 @@ variables [set_like σR R] [set_like σS S] [subsemiring_class σR R] [subsemiri open subsemiring /-- Restriction of a ring homomorphism to a subsemiring of the domain. -/ -def restrict (f : R →+* S) (s : σR) : s →+* S := f.comp $ subsemiring_class.subtype s +def dom_restrict (f : R →+* S) (s : σR) : s →+* S := f.comp $ subsemiring_class.subtype s -@[simp] lemma restrict_apply (f : R →+* S) {s : σR} (x : s) : f.restrict s x = f x := rfl +@[simp] lemma restrict_apply (f : R →+* S) {s : σR} (x : s) : f.dom_restrict s x = f x := rfl /-- Restriction of a ring homomorphism to a subsemiring of the codomain. -/ def cod_restrict (f : R →+* S) (s : σS) (h : ∀ x, f x ∈ s) : R →+* s := @@ -891,6 +891,17 @@ def cod_restrict (f : R →+* S) (s : σS) (h : ∀ x, f x ∈ s) : R →+* s := .. (f : R →* S).cod_restrict s h, .. (f : R →+ S).cod_restrict s h } +/-- The ring homomorphism from the preimage of `s` to `s`. -/ +def restrict (f : R →+* S) (s' : σR) (s : σS) (h : ∀ x ∈ s', f x ∈ s) : + s' →+* s := (f.dom_restrict s').cod_restrict s (λ x, h x x.2) + +@[simp] lemma coe_restrict_apply (f : R →+* S) (s' : σR) (s : σS) (h : ∀ x ∈ s', f x ∈ s) (x : s') : + (f.restrict s' s h x : S) = f x := rfl + +@[simp] lemma comp_restrict (f : R →+* S) (s' : σR) (s : σS) (h : ∀ x ∈ s', f x ∈ s) : + (subsemiring_class.subtype s).comp (f.restrict s' s h) = f.comp (subsemiring_class.subtype s') := +rfl + /-- Restriction of a ring homomorphism to its range interpreted as a subsemiring. This is the bundled version of `set.range_factorization`. -/ diff --git a/src/ring_theory/valuation/valuation_subring.lean b/src/ring_theory/valuation/valuation_subring.lean index 5c5257f1ee706..c66036bcc8089 100644 --- a/src/ring_theory/valuation/valuation_subring.lean +++ b/src/ring_theory/valuation/valuation_subring.lean @@ -50,6 +50,13 @@ lemma neg_mem (x : K) : x ∈ A → (-x) ∈ A := A.to_subring.neg_mem lemma mem_or_inv_mem (x : K) : x ∈ A ∨ x⁻¹ ∈ A := A.mem_or_inv_mem' _ +instance : subring_class (valuation_subring K) K := +{ zero_mem := zero_mem, + add_mem := add_mem, + one_mem := one_mem, + mul_mem := mul_mem, + neg_mem := neg_mem } + lemma to_subring_injective : function.injective (to_subring : valuation_subring K → subring K) := λ x y h, by { cases x, cases y, congr' } @@ -222,7 +229,8 @@ instance of_prime_scalar_tower (A : valuation_subring K) (P : ideal A) [P.is_pri instance of_prime_localization (A : valuation_subring K) (P : ideal A) [P.is_prime] : is_localization.at_prime (A.of_prime P) P := -by apply localization.subalgebra.is_localization_of_field K +by apply localization.subalgebra.is_localization_of_field K P.prime_compl + P.prime_compl_le_non_zero_divisors lemma le_of_prime (A : valuation_subring K) (P : ideal A) [P.is_prime] : A ≤ of_prime A P := @@ -566,8 +574,7 @@ lemma coe_mem_principal_unit_group_iff {x : A.unit_group} : A.unit_group_mul_equiv x ∈ (units.map (local_ring.residue A).to_monoid_hom).ker := begin rw [monoid_hom.mem_ker, units.ext_iff], - dsimp, - let π := ideal.quotient.mk (local_ring.maximal_ideal A), change _ ↔ π _ = _, + let π := ideal.quotient.mk (local_ring.maximal_ideal A), convert_to _ ↔ π _ = 1, rw [← π.map_one, ← sub_eq_zero, ← π.map_sub, ideal.quotient.eq_zero_iff_mem, valuation_lt_one_iff], simpa, @@ -716,6 +723,28 @@ set.subset_set_smul_iff end pointwise_actions +section + +variables {L J: Type*} [field L] [field J] + +/-- The pullback of a valuation subring `A` along a ring homomorphism `K →+* L`. -/ +def comap (A : valuation_subring L) (f : K →+* L) : + valuation_subring K := +{ mem_or_inv_mem' := λ k, by simp [valuation_subring.mem_or_inv_mem], + ..(A.to_subring.comap f) } + +@[simp] +lemma coe_comap (A : valuation_subring L) (f : K →+* L) : (A.comap f : set K) = f ⁻¹' A := rfl + +@[simp] +lemma mem_comap {A : valuation_subring L} {f : K →+* L} {x : K} : x ∈ A.comap f ↔ f x ∈ A := iff.rfl + +lemma comap_comap (A : valuation_subring J) (g : L →+* J) (f : K →+* L) : + (A.comap g).comap f = A.comap (g.comp f) := +rfl + +end + end valuation_subring namespace valuation From c55911f6166b348e1c72b08c8664e3af5f1ce334 Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Tue, 13 Dec 2022 03:30:55 +0000 Subject: [PATCH 0026/1291] chore(order/atoms): split out finiteness results (#17866) --- src/group_theory/subgroup/basic.lean | 1 + src/group_theory/sylow.lean | 1 + src/order/atoms.lean | 60 +------------------ src/order/atoms/finite.lean | 83 +++++++++++++++++++++++++++ src/order/partition/finpartition.lean | 2 +- 5 files changed, 87 insertions(+), 60 deletions(-) create mode 100644 src/order/atoms/finite.lean diff --git a/src/group_theory/subgroup/basic.lean b/src/group_theory/subgroup/basic.lean index 23362e9ec2150..bf6bb5597f37e 100644 --- a/src/group_theory/subgroup/basic.lean +++ b/src/group_theory/subgroup/basic.lean @@ -7,6 +7,7 @@ import algebra.group.conj import algebra.module.basic import algebra.order.group.inj_surj import data.countable.basic +import data.set.finite import group_theory.submonoid.centralizer import group_theory.submonoid.membership import logic.encodable.basic diff --git a/src/group_theory/sylow.lean b/src/group_theory/sylow.lean index cb6ae0218a3c8..861713c4025c6 100644 --- a/src/group_theory/sylow.lean +++ b/src/group_theory/sylow.lean @@ -9,6 +9,7 @@ import data.set_like.fintype import group_theory.group_action.conj_act import group_theory.p_group import group_theory.noncomm_pi_coprod +import order.atoms.finite /-! # Sylow theorems diff --git a/src/order/atoms.lean b/src/order/atoms.lean index 2648f19ace44e..efea364ba2fc6 100644 --- a/src/order/atoms.lean +++ b/src/order/atoms.lean @@ -3,8 +3,8 @@ Copyright (c) 2020 Aaron Anderson. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Aaron Anderson -/ -import data.set.finite import order.modular_lattice +import order.well_founded /-! # Atoms, Coatoms, and Simple Lattices @@ -42,8 +42,6 @@ which are lattices with only two elements, and related ideas. * `is_compl.is_atom_iff_is_coatom` and `is_compl.is_coatom_if_is_atom`: In a modular bounded lattice, a complement of an atom is a coatom and vice versa. * `is_atomic_iff_is_coatomic`: A modular complemented lattice is atomic iff it is coatomic. - * `finite.to_is_atomic`, `finite.to_is_coatomic`: Finite partial orders with bottom resp. top - are atomic resp. coatomic. -/ @@ -458,12 +456,6 @@ def order_iso_bool : α ≃o bool := end, ..equiv_bool } -/- It is important that `is_simple_order` is the last type-class argument of this instance, -so that type-class inference fails quickly if it doesn't apply. -/ -@[priority 200] -instance {α} [decidable_eq α] [has_le α] [bounded_order α] [is_simple_order α] : fintype α := -fintype.of_equiv bool equiv_bool.symm - /-- A simple `bounded_order` is also a `boolean_algebra`. -/ protected def boolean_algebra {α} [decidable_eq α] [lattice α] [bounded_order α] [is_simple_order α] : boolean_algebra α := @@ -535,34 +527,6 @@ instance : is_atomistic α := instance : is_coatomistic α := is_atomistic_dual_iff_is_coatomistic.1 is_simple_order.is_atomistic end is_simple_order -namespace fintype -namespace is_simple_order -variables [partial_order α] [bounded_order α] [is_simple_order α] [decidable_eq α] - -lemma univ : (finset.univ : finset α) = {⊤, ⊥} := -begin - change finset.map _ (finset.univ : finset bool) = _, - rw fintype.univ_bool, - simp only [finset.map_insert, function.embedding.coe_fn_mk, finset.map_singleton], - refl, -end - -lemma card : fintype.card α = 2 := -(fintype.of_equiv_card _).trans fintype.card_bool - -end is_simple_order -end fintype - -namespace bool - -instance : is_simple_order bool := -⟨λ a, begin - rw [← finset.mem_singleton, or.comm, ← finset.mem_insert, - top_eq_tt, bot_eq_ff, ← fintype.univ_bool], - apply finset.mem_univ, -end⟩ - -end bool theorem is_simple_order_iff_is_atom_top [partial_order α] [bounded_order α] : is_simple_order α ↔ is_atom (⊤ : α) := @@ -755,28 +719,6 @@ theorem is_atomic_iff_is_coatomic : is_atomic α ↔ is_coatomic α := end is_modular_lattice -section fintype - -open finset - -@[priority 100] -- see Note [lower instance priority] -instance finite.to_is_coatomic [partial_order α] [order_top α] [finite α] : is_coatomic α := -begin - refine is_coatomic.mk (λ b, or_iff_not_imp_left.2 (λ ht, _)), - obtain ⟨c, hc, hmax⟩ := set.finite.exists_maximal_wrt id { x : α | b ≤ x ∧ x ≠ ⊤ } - (set.to_finite _) ⟨b, le_rfl, ht⟩, - refine ⟨c, ⟨hc.2, λ y hcy, _⟩, hc.1⟩, - by_contra hyt, - obtain rfl : c = y := hmax y ⟨hc.1.trans hcy.le, hyt⟩ hcy.le, - exact (lt_self_iff_false _).mp hcy -end - -@[priority 100] -- see Note [lower instance priority] -instance finite.to_is_atomic [partial_order α] [order_bot α] [finite α] : is_atomic α := -is_coatomic_dual_iff_is_atomic.mp finite.to_is_coatomic - -end fintype - namespace set lemma is_atom_singleton (x : α) : is_atom ({x} : set α) := diff --git a/src/order/atoms/finite.lean b/src/order/atoms/finite.lean new file mode 100644 index 0000000000000..2a19d6b129053 --- /dev/null +++ b/src/order/atoms/finite.lean @@ -0,0 +1,83 @@ +/- +Copyright (c) 2020 Aaron Anderson. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Aaron Anderson +-/ +import data.set.finite +import order.atoms + +/-! +# Atoms, Coatoms, Simple Lattices, and Finiteness + +This module contains some results on atoms and simple lattices in the finite context. + +## Main results + * `finite.to_is_atomic`, `finite.to_is_coatomic`: Finite partial orders with bottom resp. top + are atomic resp. coatomic. + +-/ + +variables {α β : Type*} + +namespace is_simple_order +section decidable_eq + +/- It is important that `is_simple_order` is the last type-class argument of this instance, +so that type-class inference fails quickly if it doesn't apply. -/ +@[priority 200] +instance {α} [decidable_eq α] [has_le α] [bounded_order α] [is_simple_order α] : fintype α := +fintype.of_equiv bool equiv_bool.symm + +end decidable_eq +end is_simple_order + +namespace fintype +namespace is_simple_order +variables [partial_order α] [bounded_order α] [is_simple_order α] [decidable_eq α] + +lemma univ : (finset.univ : finset α) = {⊤, ⊥} := +begin + change finset.map _ (finset.univ : finset bool) = _, + rw fintype.univ_bool, + simp only [finset.map_insert, function.embedding.coe_fn_mk, finset.map_singleton], + refl, +end + +lemma card : fintype.card α = 2 := +(fintype.of_equiv_card _).trans fintype.card_bool + +end is_simple_order +end fintype + +namespace bool + +instance : is_simple_order bool := +⟨λ a, begin + rw [← finset.mem_singleton, or.comm, ← finset.mem_insert, + top_eq_tt, bot_eq_ff, ← fintype.univ_bool], + apply finset.mem_univ, +end⟩ + +end bool + +section fintype + +open finset + +@[priority 100] -- see Note [lower instance priority] +instance finite.to_is_coatomic [partial_order α] [order_top α] [finite α] : is_coatomic α := +begin + refine is_coatomic.mk (λ b, or_iff_not_imp_left.2 (λ ht, _)), + obtain ⟨c, hc, hmax⟩ := set.finite.exists_maximal_wrt id { x : α | b ≤ x ∧ x ≠ ⊤ } + (set.to_finite _) ⟨b, le_rfl, ht⟩, + refine ⟨c, ⟨hc.2, λ y hcy, _⟩, hc.1⟩, + by_contra hyt, + obtain rfl : c = y := hmax y ⟨hc.1.trans hcy.le, hyt⟩ hcy.le, + exact (lt_self_iff_false _).mp hcy +end + +@[priority 100] -- see Note [lower instance priority] +instance finite.to_is_atomic [partial_order α] [order_bot α] [finite α] : is_atomic α := +is_coatomic_dual_iff_is_atomic.mp finite.to_is_coatomic + +end fintype diff --git a/src/order/partition/finpartition.lean b/src/order/partition/finpartition.lean index 6a26c426fca3e..c20d35dafaaaf 100644 --- a/src/order/partition/finpartition.lean +++ b/src/order/partition/finpartition.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies, Bhavik Mehta -/ import algebra.big_operators.basic -import order.atoms +import order.atoms.finite import order.sup_indep /-! From dd4c2044a9dee231309637e7fd4e61aea1506e33 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 13 Dec 2022 03:30:56 +0000 Subject: [PATCH 0027/1291] feat(algebra/group_power/order): add monotonicity lemmas (#17895) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The existing `strict_mono_pow` lemma is renamed to `pow_strict_mono_right` to match the new `pow_strict_mono_right'`. From [this zulip thread](https://leanprover.zulipchat.com/#narrow/stream/217875-Is-there-code-for-X.3F/topic/If.20.60a.20.2B.20b.20.E2.89.A4.202.20*.20c.60.20then.20one.20of.20them.20is.20.60.E2.89.A4.20.20c.60/near/315126357). Co-authored-by: Yaël Dillies Co-authored-by: Yaël Dillies --- src/algebra/group_power/order.lean | 79 +++++++++++++++++-- src/algebra/order/monoid/lemmas.lean | 9 +++ src/number_theory/padics/padic_integers.lean | 2 +- .../polynomial/cyclotomic/basic.lean | 2 +- 4 files changed, 85 insertions(+), 7 deletions(-) diff --git a/src/algebra/group_power/order.lean b/src/algebra/group_power/order.lean index 95e0b4bc817bb..cd935027fefcb 100644 --- a/src/algebra/group_power/order.lean +++ b/src/algebra/group_power/order.lean @@ -17,7 +17,7 @@ depend on this file. open function -variables {A G M R : Type*} +variables {β A G M R : Type*} section monoid variable [monoid M] @@ -111,6 +111,37 @@ lemma right.pow_le_one_of_le (hx : x ≤ 1) : ∀ {n : ℕ}, x^n ≤ 1 end right +section covariant_lt_swap +variables [preorder β] [covariant_class M M (*) (<)] [covariant_class M M (swap (*)) (<)] + {f : β → M} + +@[to_additive strict_mono.nsmul_left] +lemma strict_mono.pow_right' (hf : strict_mono f) : ∀ {n : ℕ}, n ≠ 0 → strict_mono (λ a, f a ^ n) +| 0 hn := (hn rfl).elim +| 1 hn := by simpa +| (nat.succ $ nat.succ n) hn := + by { simp_rw pow_succ _ (n + 1), exact hf.mul' (strict_mono.pow_right' n.succ_ne_zero) } + +/-- See also `pow_strict_mono_right` -/ +@[nolint to_additive_doc, to_additive nsmul_strict_mono_left] +lemma pow_strict_mono_right' {n : ℕ} (hn : n ≠ 0) : strict_mono (λ a : M, a ^ n) := +strict_mono_id.pow_right' hn + +end covariant_lt_swap + +section covariant_le_swap +variables [preorder β] [covariant_class M M (*) (≤)] [covariant_class M M (swap (*)) (≤)] + +@[to_additive monotone.nsmul_left] +lemma monotone.pow_right {f : β → M} (hf : monotone f) : ∀ n : ℕ, monotone (λ a, f a ^ n) +| 0 := by simpa using monotone_const +| (n + 1) := by { simp_rw pow_succ, exact hf.mul' (monotone.pow_right _) } + +@[to_additive nsmul_mono_left] +lemma pow_mono_right (n : ℕ) : monotone (λ a : M, a ^ n) := monotone_id.pow_right _ + +end covariant_le_swap + @[to_additive left.pow_neg] lemma left.pow_lt_one_of_lt [covariant_class M M (*) (<)] {n : ℕ} {x : M} (hn : 0 < n) (h : x < 1) : x^n < 1 := @@ -162,6 +193,44 @@ lemma pow_lt_pow_iff' (ha : 1 < a) : a ^ m < a ^ n ↔ m < n := (pow_strict_mono end covariant_le +section covariant_le_swap +variables [covariant_class M M (*) (≤)] [covariant_class M M (swap (*)) (≤)] + +@[to_additive lt_of_nsmul_lt_nsmul] +lemma lt_of_pow_lt_pow' {a b : M} (n : ℕ) : a ^ n < b ^ n → a < b := (pow_mono_right _).reflect_lt + +@[to_additive] +lemma min_lt_max_of_mul_lt_mul {a b c d : M} (h : a * b < c * d) : min a b < max c d := +lt_of_pow_lt_pow' 2 $ by { simp_rw pow_two, exact (mul_le_mul' inf_le_left + inf_le_right).trans_lt (h.trans_le $ mul_le_mul' le_sup_left le_sup_right) } + +@[to_additive min_lt_of_add_lt_two_nsmul] +lemma min_lt_of_mul_lt_sq {a b c : M} (h : a * b < c ^ 2) : min a b < c := +by simpa using min_lt_max_of_mul_lt_mul (h.trans_eq $ pow_two _) + +@[to_additive lt_max_of_two_nsmul_lt_add] +lemma lt_max_of_sq_lt_mul {a b c : M} (h : a ^ 2 < b * c) : a < max b c := +by simpa using min_lt_max_of_mul_lt_mul ((pow_two _).symm.trans_lt h) + +end covariant_le_swap + +section covariant_lt_swap +variables [covariant_class M M (*) (<)] [covariant_class M M (swap (*)) (<)] + +@[to_additive le_of_nsmul_le_nsmul] +lemma le_of_pow_le_pow' {a b : M} {n : ℕ} (hn : n ≠ 0) : a ^ n ≤ b ^ n → a ≤ b := +(pow_strict_mono_right' hn).le_iff_le.1 + +@[to_additive min_le_of_add_le_two_nsmul] +lemma min_le_of_mul_le_sq {a b c : M} (h : a * b ≤ c ^ 2) : min a b ≤ c := +by simpa using min_le_max_of_mul_le_mul (h.trans_eq $ pow_two _) + +@[to_additive le_max_of_two_nsmul_le_add] +lemma le_max_of_sq_le_mul {a b c : M} (h : a ^ 2 ≤ b * c) : a ≤ max b c := +by simpa using min_le_max_of_mul_le_mul ((pow_two _).symm.trans_le h) + +end covariant_lt_swap + @[to_additive left.nsmul_neg_iff] lemma left.pow_lt_one_iff [covariant_class M M (*) (<)] {n : ℕ} {x : M} (hn : 0 < n) : x^n < 1 ↔ x < 1 := @@ -271,19 +340,19 @@ lemma pow_lt_pow_of_lt_left (h : x < y) (hx : 0 ≤ x) : ∀ {n : ℕ}, 0 < n lemma strict_mono_on_pow (hn : 0 < n) : strict_mono_on (λ x : R, x ^ n) (set.Ici 0) := λ x hx y hy h, pow_lt_pow_of_lt_left h hx hn -lemma strict_mono_pow (h : 1 < a) : strict_mono (λ n : ℕ, a ^ n) := +lemma pow_strict_mono_right (h : 1 < a) : strict_mono (λ n : ℕ, a ^ n) := have 0 < a := zero_le_one.trans_lt h, strict_mono_nat_of_lt_succ $ λ n, by simpa only [one_mul, pow_succ] using mul_lt_mul h (le_refl (a ^ n)) (pow_pos this _) this.le lemma pow_lt_pow (h : 1 < a) (h2 : n < m) : a ^ n < a ^ m := -strict_mono_pow h h2 +pow_strict_mono_right h h2 lemma pow_lt_pow_iff (h : 1 < a) : a ^ n < a ^ m ↔ n < m := -(strict_mono_pow h).lt_iff_lt +(pow_strict_mono_right h).lt_iff_lt lemma pow_le_pow_iff (h : 1 < a) : a ^ n ≤ a ^ m ↔ n ≤ m := -(strict_mono_pow h).le_iff_le +(pow_strict_mono_right h).le_iff_le lemma strict_anti_pow (h₀ : 0 < a) (h₁ : a < 1) : strict_anti (λ n : ℕ, a ^ n) := strict_anti_nat_of_succ_lt $ λ n, diff --git a/src/algebra/order/monoid/lemmas.lean b/src/algebra/order/monoid/lemmas.lean index c859da1aaeb10..03107416821fb 100644 --- a/src/algebra/order/monoid/lemmas.lean +++ b/src/algebra/order/monoid/lemmas.lean @@ -5,6 +5,7 @@ Authors: Jeremy Avigad, Leonardo de Moura, Mario Carneiro, Johannes Hölzl, Dami Yuyang Zhao -/ import algebra.covariant_and_contravariant +import order.min_max /-! # Ordered monoids @@ -237,6 +238,14 @@ le_antisymm (le_of_mul_le_mul_right' h.le) (le_of_mul_le_mul_right' h.ge) end partial_order +section linear_order +variables [linear_order α] {a b c d : α} [covariant_class α α (*) (<)] + [covariant_class α α (swap (*)) (<)] + +@[to_additive] lemma min_le_max_of_mul_le_mul (h : a * b ≤ c * d) : min a b ≤ max c d := +by { simp_rw [min_le_iff, le_max_iff], contrapose! h, exact mul_lt_mul_of_lt_of_lt h.1.1 h.2.2 } + +end linear_order end has_mul -- using one diff --git a/src/number_theory/padics/padic_integers.lean b/src/number_theory/padics/padic_integers.lean index f6d3a1373ac73..5b025a364f284 100644 --- a/src/number_theory/padics/padic_integers.lean +++ b/src/number_theory/padics/padic_integers.lean @@ -421,7 +421,7 @@ begin have aux : ∀ n : ℕ, 0 < (p ^ n : ℝ), { apply pow_pos, exact_mod_cast hp.1.pos }, rw [inv_le_inv (aux _) (aux _)], - have : p ^ n ≤ p ^ k ↔ n ≤ k := (strict_mono_pow hp.1.one_lt).le_iff_le, + have : p ^ n ≤ p ^ k ↔ n ≤ k := (pow_strict_mono_right hp.1.one_lt).le_iff_le, rw [← this], norm_cast end diff --git a/src/ring_theory/polynomial/cyclotomic/basic.lean b/src/ring_theory/polynomial/cyclotomic/basic.lean index 502ba0b34b77b..836c87d3708fa 100644 --- a/src/ring_theory/polynomial/cyclotomic/basic.lean +++ b/src/ring_theory/polynomial/cyclotomic/basic.lean @@ -1014,7 +1014,7 @@ begin rw [cyclotomic_mul_prime_pow_eq R (ne_zero.not_char_dvd R p m) hk, is_root.def, eval_pow, h, zero_pow], simp only [tsub_pos_iff_lt], - apply strict_mono_pow hp.out.one_lt (nat.pred_lt hk.ne') } + apply pow_strict_mono_right hp.out.one_lt (nat.pred_lt hk.ne') } end end char_p From ec4bdab82c3c8fc4befc4aace5455fc5753177ec Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 13 Dec 2022 03:30:57 +0000 Subject: [PATCH 0028/1291] chore(group_theory/perm/cycle/basic): Move declarations around (#17906) Move the `same_cycle` lemmas before the `is_cycle` ones in order to reduce the diff in #17898. Co-authored-by: Eric Wieser --- src/group_theory/perm/cycle/basic.lean | 209 ++++++++++++------------- 1 file changed, 103 insertions(+), 106 deletions(-) diff --git a/src/group_theory/perm/cycle/basic.lean b/src/group_theory/perm/cycle/basic.lean index 8b6811ff8d37e..fd0239172c1b5 100644 --- a/src/group_theory/perm/cycle/basic.lean +++ b/src/group_theory/perm/cycle/basic.lean @@ -14,9 +14,9 @@ import logic.equiv.fintype In the following, `f : equiv.perm β`. +* `equiv.perm.same_cycle`: `f.same_cycle x y` when `x` and `y` are in the same cycle of `f`. * `equiv.perm.is_cycle`: `f.is_cycle` when two nonfixed points of `β` are related by repeated application of `f`. -* `equiv.perm.same_cycle`: `f.same_cycle x y` when `x` and `y` are in the same cycle of `f`. The following two definitions require that `β` is a `fintype`: @@ -40,15 +40,99 @@ The following two definitions require that `β` is a `fintype`: namespace equiv.perm open equiv function finset -variables {α : Type*} {β : Type*} [decidable_eq α] +variables {α β : Type*} + +/-! ### `same_cycle` -/ + +section same_cycle +variables {f g : perm α} {p : α → Prop} {x y z : α} + +/-- The equivalence relation indicating that two points are in the same cycle of a permutation. -/ +def same_cycle (f : perm β) (x y : β) : Prop := ∃ i : ℤ, (f ^ i) x = y + +@[refl] lemma same_cycle.refl (f : perm β) (x : β) : same_cycle f x x := ⟨0, rfl⟩ + +@[symm] lemma same_cycle.symm : same_cycle f x y → same_cycle f y x := +λ ⟨i, hi⟩, ⟨-i, by rw [zpow_neg, ← hi, inv_apply_self]⟩ + +@[trans] lemma same_cycle.trans : same_cycle f x y → same_cycle f y z → same_cycle f x z := +λ ⟨i, hi⟩ ⟨j, hj⟩, ⟨j + i, by rw [zpow_add, mul_apply, hi, hj]⟩ + +lemma same_cycle_inv (f : perm α) : same_cycle f⁻¹ x y ↔ same_cycle f x y := +(equiv.neg _).exists_congr_left.trans $ by simp [same_cycle] + +lemma same_cycle.apply_eq_self_iff : same_cycle f x y → (f x = x ↔ f y = y) := +λ ⟨i, hi⟩, by rw [← hi, ← mul_apply, ← zpow_one_add, add_comm, zpow_add_one, mul_apply, + (f ^ i).injective.eq_iff] + +lemma same_cycle_apply : same_cycle f x (f y) ↔ same_cycle f x y := +⟨λ ⟨i, hi⟩, ⟨-1 + i, by rw [zpow_add, mul_apply, hi, zpow_neg_one, inv_apply_self]⟩, + λ ⟨i, hi⟩, ⟨1 + i, by rw [zpow_add, mul_apply, hi, zpow_one]⟩⟩ + +lemma same_cycle_inv_apply : same_cycle f x (f⁻¹ y) ↔ same_cycle f x y := +by rw [← same_cycle_inv, same_cycle_apply, same_cycle_inv] + +@[simp] lemma same_cycle_pow_left_iff {n : ℕ} : same_cycle f ((f ^ n) x) y ↔ same_cycle f x y := +begin + split, + { rintro ⟨k, rfl⟩, + use (k + n), + simp [zpow_add] }, + { rintro ⟨k, rfl⟩, + use (k - n), + rw [←zpow_coe_nat, ←mul_apply, ←zpow_add, int.sub_add_cancel] } +end + +@[simp] lemma same_cycle_zpow_left_iff {n : ℤ} : same_cycle f ((f ^ n) x) y ↔ same_cycle f x y := +begin + cases n, + { exact same_cycle_pow_left_iff }, + { rw [zpow_neg_succ_of_nat, ←inv_pow, ←same_cycle_inv, same_cycle_pow_left_iff, same_cycle_inv] } +end + +lemma same_cycle.nat' [finite α] : same_cycle f x y → ∃ i < order_of f, (f ^ i) x = y := +begin + classical, + rintro ⟨k, rfl⟩, + use (k % order_of f).nat_abs, + have h₀ := int.coe_nat_pos.mpr (order_of_pos f), + have h₁ := int.mod_nonneg k h₀.ne', + rw [←zpow_coe_nat, int.nat_abs_of_nonneg h₁, ←zpow_eq_mod_order_of], + refine ⟨_, rfl⟩, + rw [←int.coe_nat_lt, int.nat_abs_of_nonneg h₁], + exact int.mod_lt_of_pos _ h₀, +end + +lemma same_cycle.nat'' [finite α] (h : same_cycle f x y) : + ∃ (i : ℕ) (hpos : 0 < i) (h : i ≤ order_of f), (f ^ i) x = y := +begin + classical, + obtain ⟨_|i, hi, rfl⟩ := h.nat', + { refine ⟨order_of f, order_of_pos f, le_rfl, _⟩, + rw [pow_order_of_eq_one, pow_zero] }, + { exact ⟨i.succ, i.zero_lt_succ, hi.le, rfl⟩ } +end + +instance [fintype α] [decidable_eq α] (f : perm α) : decidable_rel (same_cycle f) := +λ x y, decidable_of_iff (∃ n ∈ list.range (fintype.card (perm α)), (f ^ n) x = y) +⟨λ ⟨n, _, hn⟩, ⟨n, hn⟩, λ ⟨i, hi⟩, ⟨(i % order_of f).nat_abs, list.mem_range.2 + (int.coe_nat_lt.1 $ + by { rw int.nat_abs_of_nonneg (int.mod_nonneg _ $ int.coe_nat_ne_zero.2 (order_of_pos _).ne'), + { refine (int.mod_lt _ $ int.coe_nat_ne_zero_iff_pos.2 $ order_of_pos _).trans_le _, + simp [order_of_le_card_univ] }, + apply_instance }), + by { rw [← zpow_coe_nat, int.nat_abs_of_nonneg (int.mod_nonneg _ $ + int.coe_nat_ne_zero_iff_pos.2 $ order_of_pos _), ← zpow_eq_mod_order_of, hi], + apply_instance }⟩⟩ -section sign_cycle +end same_cycle /-! ### `is_cycle` -/ -variables [fintype α] +section is_cycle +variables [fintype α] [decidable_eq α] /-- A permutation is a cycle when any two nonfixed points of the permutation are related by repeated application of the permutation. -/ @@ -102,6 +186,10 @@ let ⟨a, ha⟩ := hg.2 x hx in let ⟨b, hb⟩ := hg.2 y hy in ⟨b - a, by rw [← ha, ← mul_apply, ← zpow_add, sub_add_cancel, hb]⟩ +lemma is_cycle.same_cycle {f : perm β} (hf : is_cycle f) {x y : β} (hx : f x ≠ x) (hy : f y ≠ y) : + same_cycle f x y := +hf.exists_zpow_eq hx hy + lemma is_cycle.exists_pow_eq [finite β] {f : perm β} (hf : is_cycle f) {x y : β} (hx : f x ≠ x) (hy : f y ≠ y) : ∃ i : ℕ, (f ^ i) x = y := let ⟨n, hn⟩ := hf.exists_zpow_eq hx hy in @@ -309,78 +397,6 @@ begin { exact (hb (extend_domain_apply_not_subtype _ _ pb)).elim } end -lemma nodup_of_pairwise_disjoint_cycles {l : list (perm β)} (h1 : ∀ f ∈ l, is_cycle f) - (h2 : l.pairwise disjoint) : l.nodup := -nodup_of_pairwise_disjoint (λ h, (h1 1 h).ne_one rfl) h2 - -end sign_cycle - -/-! -### `same_cycle` --/ - -/-- The equivalence relation indicating that two points are in the same cycle of a permutation. -/ -def same_cycle (f : perm β) (x y : β) : Prop := ∃ i : ℤ, (f ^ i) x = y - -@[refl] lemma same_cycle.refl (f : perm β) (x : β) : same_cycle f x x := ⟨0, rfl⟩ - -@[symm] lemma same_cycle.symm {f : perm β} {x y : β} : same_cycle f x y → same_cycle f y x := -λ ⟨i, hi⟩, ⟨-i, by rw [zpow_neg, ← hi, inv_apply_self]⟩ - -@[trans] lemma same_cycle.trans {f : perm β} {x y z : β} : - same_cycle f x y → same_cycle f y z → same_cycle f x z := -λ ⟨i, hi⟩ ⟨j, hj⟩, ⟨j + i, by rw [zpow_add, mul_apply, hi, hj]⟩ - -lemma same_cycle.apply_eq_self_iff {f : perm β} {x y : β} : - same_cycle f x y → (f x = x ↔ f y = y) := -λ ⟨i, hi⟩, by rw [← hi, ← mul_apply, ← zpow_one_add, add_comm, zpow_add_one, mul_apply, - (f ^ i).injective.eq_iff] - -lemma is_cycle.same_cycle {f : perm β} (hf : is_cycle f) {x y : β} - (hx : f x ≠ x) (hy : f y ≠ y) : same_cycle f x y := -hf.exists_zpow_eq hx hy - -lemma same_cycle.nat' [finite β] {f : perm β} {x y : β} (h : same_cycle f x y) : - ∃ (i : ℕ) (h : i < order_of f), (f ^ i) x = y := -begin - classical, - obtain ⟨k, rfl⟩ := h, - use ((k % order_of f).nat_abs), - have h₀ := int.coe_nat_pos.mpr (order_of_pos f), - have h₁ := int.mod_nonneg k h₀.ne', - rw [←zpow_coe_nat, int.nat_abs_of_nonneg h₁, ←zpow_eq_mod_order_of], - refine ⟨_, rfl⟩, - rw [←int.coe_nat_lt, int.nat_abs_of_nonneg h₁], - exact int.mod_lt_of_pos _ h₀, -end - -lemma same_cycle.nat'' [finite β] {f : perm β} {x y : β} (h : same_cycle f x y) : - ∃ (i : ℕ) (hpos : 0 < i) (h : i ≤ order_of f), (f ^ i) x = y := -begin - classical, - obtain ⟨_|i, hi, rfl⟩ := h.nat', - { refine ⟨order_of f, order_of_pos f, le_rfl, _⟩, - rw [pow_order_of_eq_one, pow_zero] }, - { exact ⟨i.succ, i.zero_lt_succ, hi.le, rfl⟩ } -end - -instance [fintype α] (f : perm α) : decidable_rel (same_cycle f) := -λ x y, decidable_of_iff (∃ n ∈ list.range (fintype.card (perm α)), (f ^ n) x = y) -⟨λ ⟨n, _, hn⟩, ⟨n, hn⟩, λ ⟨i, hi⟩, ⟨(i % order_of f).nat_abs, list.mem_range.2 - (int.coe_nat_lt.1 $ - by { rw int.nat_abs_of_nonneg (int.mod_nonneg _ - (int.coe_nat_ne_zero_iff_pos.2 (order_of_pos _))), - { refine (int.mod_lt _ $ int.coe_nat_ne_zero_iff_pos.2 $ order_of_pos _).trans_le _, - simp [order_of_le_card_univ] }, - apply_instance }), - by { rw [← zpow_coe_nat, int.nat_abs_of_nonneg (int.mod_nonneg _ - (int.coe_nat_ne_zero_iff_pos.2 (order_of_pos _))), ← zpow_eq_mod_order_of, hi], - apply_instance }⟩⟩ - -lemma same_cycle_apply {f : perm β} {x y : β} : same_cycle f x (f y) ↔ same_cycle f x y := -⟨λ ⟨i, hi⟩, ⟨-1 + i, by rw [zpow_add, mul_apply, hi, zpow_neg_one, inv_apply_self]⟩, - λ ⟨i, hi⟩, ⟨1 + i, by rw [zpow_add, mul_apply, hi, zpow_one]⟩⟩ - lemma same_cycle_cycle {f : perm β} {x : β} (hx : f x ≠ x) : is_cycle f ↔ (∀ {y}, same_cycle f x y ↔ f y ≠ y) := ⟨λ hf y, ⟨λ ⟨i, hi⟩ hy, hx $ @@ -389,36 +405,13 @@ lemma same_cycle_cycle {f : perm β} {x : β} (hx : f x ≠ x) : is_cycle f ↔ hf.exists_zpow_eq hx⟩, λ h, ⟨x, hx, λ y hy, h.2 hy⟩⟩ -lemma same_cycle_inv (f : perm β) {x y : β} : same_cycle f⁻¹ x y ↔ same_cycle f x y := -⟨λ ⟨i, hi⟩, ⟨-i, by rw [zpow_neg, ← inv_zpow, hi]⟩, - λ ⟨i, hi⟩, ⟨-i, by rw [zpow_neg, ← inv_zpow, inv_inv, hi]⟩ ⟩ - -lemma same_cycle_inv_apply {f : perm β} {x y : β} : same_cycle f x (f⁻¹ y) ↔ same_cycle f x y := -by rw [← same_cycle_inv, same_cycle_apply, same_cycle_inv] - -@[simp] lemma same_cycle_pow_left_iff {f : perm β} {x y : β} {n : ℕ} : - same_cycle f ((f ^ n) x) y ↔ same_cycle f x y := -begin - split, - { rintro ⟨k, rfl⟩, - use (k + n), - simp [zpow_add] }, - { rintro ⟨k, rfl⟩, - use (k - n), - rw [←zpow_coe_nat, ←mul_apply, ←zpow_add, int.sub_add_cancel] } -end - -@[simp] lemma same_cycle_zpow_left_iff {f : perm β} {x y : β} {n : ℤ} : - same_cycle f ((f ^ n) x) y ↔ same_cycle f x y := -begin - cases n, - { exact same_cycle_pow_left_iff }, - { rw [zpow_neg_succ_of_nat, ←inv_pow, ←same_cycle_inv, same_cycle_pow_left_iff, same_cycle_inv] } -end +lemma nodup_of_pairwise_disjoint_cycles {l : list (perm β)} (h1 : ∀ f ∈ l, is_cycle f) + (h2 : l.pairwise disjoint) : l.nodup := +nodup_of_pairwise_disjoint (λ h, (h1 1 h).ne_one rfl) h2 /-- Unlike `support_congr`, which assumes that `∀ (x ∈ g.support), f x = g x)`, here we have the weaker assumption that `∀ (x ∈ f.support), f x = g x`. -/ -lemma is_cycle.support_congr [fintype α] {f g : perm α} (hf : is_cycle f) (hg : is_cycle g) +lemma is_cycle.support_congr {f g : perm α} (hf : is_cycle f) (hg : is_cycle g) (h : f.support ⊆ g.support) (h' : ∀ (x ∈ f.support), f x = g x) : f = g := begin have : f.support = g.support, @@ -439,7 +432,7 @@ end /-- If two cyclic permutations agree on all terms in their intersection, and that intersection is not empty, then the two cyclic permutations must be equal. -/ -lemma is_cycle.eq_on_support_inter_nonempty_congr [fintype α] {f g : perm α} +lemma is_cycle.eq_on_support_inter_nonempty_congr {f g : perm α} (hf : is_cycle f) (hg : is_cycle g) (h : ∀ (x ∈ f.support ∩ g.support), f x = g x) {x : α} (hx : f x = g x) (hx' : x ∈ f.support) : f = g := begin @@ -453,7 +446,7 @@ begin exact hf.support_congr hg this h end -lemma is_cycle.support_pow_eq_iff [fintype α] {f : perm α} (hf : is_cycle f) {n : ℕ} : +lemma is_cycle.support_pow_eq_iff {f : perm α} (hf : is_cycle f) {n : ℕ} : support (f ^ n) = support f ↔ ¬ order_of f ∣ n := begin rw order_of_dvd_iff_pow_eq_one, @@ -544,7 +537,7 @@ begin contradiction }} end -lemma is_cycle.mem_support_pos_pow_iff_of_lt_order_of [fintype α] {f : perm α} (hf : is_cycle f) +lemma is_cycle.mem_support_pos_pow_iff_of_lt_order_of {f : perm α} (hf : is_cycle f) {n : ℕ} (npos : 0 < n) (hn : n < order_of f) {x : α} : x ∈ (f ^ n).support ↔ x ∈ f.support := begin @@ -570,11 +563,15 @@ begin exact support_pow_le f n end +end is_cycle + /-! ### `cycle_of` -/ +variables [decidable_eq α] + /-- `f.cycle_of x` is the cycle of the permutation `f` to which `x` belongs. -/ def cycle_of [fintype α] (f : perm α) (x : α) : perm α := of_subtype (subtype_perm f (λ _, same_cycle_apply.symm) : perm {y // same_cycle f x y}) From f721cdca3b50bbe616dfc291620d51fcc06bfea4 Mon Sep 17 00:00:00 2001 From: sgouezel Date: Tue, 13 Dec 2022 06:33:12 +0000 Subject: [PATCH 0029/1291] feat(measure_theory/measure/haar_lebesgue): Lebesgue measure coming from an alternating map (#17583) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit To an `n`-alternating form `ω` on an `n`-dimensional space, we associate a Lebesgue measure `ω.measure`. We show that it satisfies `ω.measure (parallelepiped v) = |ω v|` when `v` is a family of `n` vectors, where `parallelepiped v` is the parallelepiped spanned by these vectors. In the special case of inner product spaces, this gives a canonical measure when one takes for the alternating form the canonical one associated to some orientation. Therefore, we register a `measure_space` instance on finite-dimensional inner product spaces. As the real line is a special case of inner product space, we need to redefine the Lebesgue measure there to avoid having two competing `measure_space` instances. --- src/analysis/inner_product_space/pi_L2.lean | 14 +- src/analysis/normed/group/basic.lean | 5 + src/linear_algebra/determinant.lean | 4 + src/linear_algebra/finite_dimensional.lean | 9 + src/measure_theory/group/measure.lean | 22 ++- src/measure_theory/measure/haar_lebesgue.lean | 82 ++++++--- src/measure_theory/measure/haar_of_basis.lean | 160 ++++++++++++++++++ src/measure_theory/measure/haar_of_inner.lean | 69 ++++++++ src/measure_theory/measure/lebesgue.lean | 83 ++++----- src/measure_theory/measure/measure_space.lean | 4 + 10 files changed, 384 insertions(+), 68 deletions(-) create mode 100644 src/measure_theory/measure/haar_of_basis.lean create mode 100644 src/measure_theory/measure/haar_of_inner.lean diff --git a/src/analysis/inner_product_space/pi_L2.lean b/src/analysis/inner_product_space/pi_L2.lean index e66ef1bc6a3c0..ab8caa6dc83db 100644 --- a/src/analysis/inner_product_space/pi_L2.lean +++ b/src/analysis/inner_product_space/pi_L2.lean @@ -656,13 +656,25 @@ let ⟨w, hw, hw', hw''⟩ := (orthonormal_empty 𝕜 E).exists_orthonormal_basi ⟨w, hw, hw''⟩ /-- A finite-dimensional `inner_product_space` has an orthonormal basis. -/ -def std_orthonormal_basis : orthonormal_basis (fin (finrank 𝕜 E)) 𝕜 E := +@[irreducible] def std_orthonormal_basis : orthonormal_basis (fin (finrank 𝕜 E)) 𝕜 E := begin let b := classical.some (classical.some_spec $ exists_orthonormal_basis 𝕜 E), rw [finrank_eq_card_basis b.to_basis], exact b.reindex (fintype.equiv_fin_of_card_eq rfl), end +/-- An orthonormal basis of `ℝ` is made either of the vector `1`, or of the vector `-1`. -/ +lemma orthonormal_basis_one_dim (b : orthonormal_basis ι ℝ ℝ) : + ⇑b = (λ _, (1 : ℝ)) ∨ ⇑b = (λ _, (-1 : ℝ)) := +begin + haveI : unique ι, from b.to_basis.unique, + have : b default = 1 ∨ b default = - 1, + { have : ‖b default‖ = 1, from b.orthonormal.1 _, + rwa [real.norm_eq_abs, abs_eq (zero_le_one : (0 : ℝ) ≤ 1)] at this }, + rw eq_const_of_unique b, + refine this.imp _ _; simp, +end + variables {𝕜 E} section subordinate_orthonormal_basis diff --git a/src/analysis/normed/group/basic.lean b/src/analysis/normed/group/basic.lean index 88fc629764a1c..abbeb4e1b952e 100644 --- a/src/analysis/normed/group/basic.lean +++ b/src/analysis/normed/group/basic.lean @@ -1199,9 +1199,14 @@ lemma le_norm_self (r : ℝ) : r ≤ ‖r‖ := le_abs_self r lemma nnnorm_of_nonneg (hr : 0 ≤ r) : ‖r‖₊ = ⟨r, hr⟩ := nnreal.eq $ norm_of_nonneg hr +@[simp] lemma nnnorm_abs (r : ℝ) : ‖(|r|)‖₊ = ‖r‖₊ := by simp [nnnorm] + lemma ennnorm_eq_of_real (hr : 0 ≤ r) : (‖r‖₊ : ℝ≥0∞) = ennreal.of_real r := by { rw [← of_real_norm_eq_coe_nnnorm, norm_of_nonneg hr] } +lemma ennnorm_eq_of_real_abs (r : ℝ) : (‖r‖₊ : ℝ≥0∞) = ennreal.of_real (|r|) := +by rw [← real.nnnorm_abs r, real.ennnorm_eq_of_real (abs_nonneg _)] + lemma to_nnreal_eq_nnnorm_of_nonneg (hr : 0 ≤ r) : r.to_nnreal = ‖r‖₊ := begin rw real.to_nnreal_of_nonneg hr, diff --git a/src/linear_algebra/determinant.lean b/src/linear_algebra/determinant.lean index 832a3770de7ce..b670bd9a2e843 100644 --- a/src/linear_algebra/determinant.lean +++ b/src/linear_algebra/determinant.lean @@ -208,6 +208,10 @@ by simp [← to_matrix_eq_to_matrix'] linear_map.det (matrix.to_lin b b f) = f.det := by rw [← linear_map.det_to_matrix b, linear_map.to_matrix_to_lin] +@[simp] lemma det_to_lin' (f : matrix ι ι R) : + linear_map.det (f.to_lin') = f.det := +by simp only [← to_lin_eq_to_lin', det_to_lin] + /-- To show `P f.det` it suffices to consider `P (to_matrix _ _ f).det` and `P 1`. -/ @[elab_as_eliminator] lemma det_cases [decidable_eq M] {P : A → Prop} (f : M →ₗ[A] M) diff --git a/src/linear_algebra/finite_dimensional.lean b/src/linear_algebra/finite_dimensional.lean index 1b98f9bb85e4c..d0ee8c0ce5983 100644 --- a/src/linear_algebra/finite_dimensional.lean +++ b/src/linear_algebra/finite_dimensional.lean @@ -209,6 +209,15 @@ begin rw [cardinal.mk_fintype, finrank_eq_card_basis h] end +/-- Given a basis of a division ring over itself indexed by a type `ι`, then `ι` is `unique`. -/ +noncomputable def basis.unique {ι : Type*} (b : basis ι K K) : unique ι := +begin + have A : cardinal.mk ι = ↑(finite_dimensional.finrank K K) := + (finite_dimensional.finrank_eq_card_basis' b).symm, + simp only [cardinal.eq_one_iff_unique, finite_dimensional.finrank_self, algebra_map.coe_one] at A, + exact nonempty.some ((unique_iff_subsingleton_and_nonempty _).2 A), +end + variables (K V) /-- A finite dimensional vector space has a basis indexed by `fin (finrank K V)`. -/ diff --git a/src/measure_theory/group/measure.lean b/src/measure_theory/group/measure.lean index f9e58e0ed76a6..b3d09a47c4246 100644 --- a/src/measure_theory/group/measure.lean +++ b/src/measure_theory/group/measure.lean @@ -27,7 +27,7 @@ We also give analogues of all these notions in the additive world. noncomputable theory -open_locale ennreal pointwise big_operators topological_space +open_locale nnreal ennreal pointwise big_operators topological_space open has_inv set function measure_theory.measure filter variables {𝕜 G H : Type*} [measurable_space G] [measurable_space H] @@ -70,14 +70,26 @@ is_mul_left_invariant.map_mul_left_eq_self g lemma map_mul_right_eq_self (μ : measure G) [is_mul_right_invariant μ] (g : G) : map (* g) μ = μ := is_mul_right_invariant.map_mul_right_eq_self g -@[to_additive] -instance [is_mul_left_invariant μ] (c : ℝ≥0∞) : is_mul_left_invariant (c • μ) := +@[to_additive measure_theory.is_add_left_invariant_smul] +instance is_mul_left_invariant_smul [is_mul_left_invariant μ] (c : ℝ≥0∞) : + is_mul_left_invariant (c • μ) := ⟨λ g, by rw [measure.map_smul, map_mul_left_eq_self]⟩ -@[to_additive] -instance [is_mul_right_invariant μ] (c : ℝ≥0∞) : is_mul_right_invariant (c • μ) := +@[to_additive measure_theory.is_add_right_invariant_smul] +instance is_mul_right_invariant_smul [is_mul_right_invariant μ] (c : ℝ≥0∞) : + is_mul_right_invariant (c • μ) := ⟨λ g, by rw [measure.map_smul, map_mul_right_eq_self]⟩ +@[to_additive measure_theory.is_add_left_invariant_smul_nnreal] +instance is_mul_left_invariant_smul_nnreal [is_mul_left_invariant μ] (c : ℝ≥0) : + is_mul_left_invariant (c • μ) := +measure_theory.is_mul_left_invariant_smul (c : ℝ≥0∞) + +@[to_additive measure_theory.is_add_right_invariant_smul_nnreal] +instance is_mul_right_invariant_smul_nnreal [is_mul_right_invariant μ] (c : ℝ≥0) : + is_mul_right_invariant (c • μ) := +measure_theory.is_mul_right_invariant_smul (c : ℝ≥0∞) + section has_measurable_mul variables [has_measurable_mul G] diff --git a/src/measure_theory/measure/haar_lebesgue.lean b/src/measure_theory/measure/haar_lebesgue.lean index 6134d647b3f90..b52c0c1122241 100644 --- a/src/measure_theory/measure/haar_lebesgue.lean +++ b/src/measure_theory/measure/haar_lebesgue.lean @@ -30,6 +30,11 @@ We deduce basic properties of any Haar measure on a finite dimensional real vect * `add_haar_closed_ball`: the measure of `closed_ball x r` is `r ^ dim * μ (ball 0 1)`. * `add_haar_sphere`: spheres have zero measure. +This makes it possible to associate a Lebesgue measure to an `n`-alternating map in dimension `n`. +This measure is called `alternating_map.measure`. Its main property is +`ω.measure_parallelepiped v`, stating that the associated measure of the parallelepiped spanned +by vectors `v₁, ..., vₙ` is given by `|ω v|`. + We also show that a Lebesgue density point `x` of a set `s` (with respect to closed balls) has density one for the rescaled copies `{x} + r • t` of a given set `t` with positive measure, in `tendsto_add_haar_inter_smul_one_of_density_one`. In particular, `s` intersects `{x} + r • t` for @@ -67,9 +72,6 @@ open measure topological_space.positive_compacts finite_dimensional lemma add_haar_measure_eq_volume : add_haar_measure Icc01 = volume := by { convert (add_haar_measure_unique volume Icc01).symm, simp [Icc01] } -instance : is_add_haar_measure (volume : measure ℝ) := -by { rw ← add_haar_measure_eq_volume, apply_instance } - /-- The Haar measure equals the Lebesgue measure on `ℝ^ι`. -/ lemma add_haar_measure_eq_volume_pi (ι : Type*) [fintype ι] : add_haar_measure (pi_Icc01 ι) = volume := @@ -201,9 +203,11 @@ begin real.map_linear_map_volume_pi_eq_smul_volume_pi hf, smul_comm], end +variables {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] [measurable_space E] + [borel_space E] [finite_dimensional ℝ E] (μ : measure E) [is_add_haar_measure μ] + {F : Type*} [normed_add_comm_group F] [normed_space ℝ F] [complete_space F] + lemma map_linear_map_add_haar_eq_smul_add_haar - {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] [measurable_space E] [borel_space E] - [finite_dimensional ℝ E] (μ : measure E) [is_add_haar_measure μ] {f : E →ₗ[ℝ] E} (hf : f.det ≠ 0) : measure.map f μ = ennreal.of_real (abs (f.det)⁻¹) • μ := begin @@ -237,8 +241,6 @@ end /-- The preimage of a set `s` under a linear map `f` with nonzero determinant has measure equal to `μ s` times the absolute value of the inverse of the determinant of `f`. -/ @[simp] lemma add_haar_preimage_linear_map - {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] [measurable_space E] [borel_space E] - [finite_dimensional ℝ E] (μ : measure E) [is_add_haar_measure μ] {f : E →ₗ[ℝ] E} (hf : f.det ≠ 0) (s : set E) : μ (f ⁻¹' s) = ennreal.of_real (abs (f.det)⁻¹) * μ s := calc μ (f ⁻¹' s) = measure.map f μ s : @@ -250,8 +252,6 @@ calc μ (f ⁻¹' s) = measure.map f μ s : /-- The preimage of a set `s` under a continuous linear map `f` with nonzero determinant has measure equal to `μ s` times the absolute value of the inverse of the determinant of `f`. -/ @[simp] lemma add_haar_preimage_continuous_linear_map - {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] [measurable_space E] [borel_space E] - [finite_dimensional ℝ E] (μ : measure E) [is_add_haar_measure μ] {f : E →L[ℝ] E} (hf : linear_map.det (f : E →ₗ[ℝ] E) ≠ 0) (s : set E) : μ (f ⁻¹' s) = ennreal.of_real (abs (linear_map.det (f : E →ₗ[ℝ] E))⁻¹) * μ s := add_haar_preimage_linear_map μ hf s @@ -259,8 +259,6 @@ add_haar_preimage_linear_map μ hf s /-- The preimage of a set `s` under a linear equiv `f` has measure equal to `μ s` times the absolute value of the inverse of the determinant of `f`. -/ @[simp] lemma add_haar_preimage_linear_equiv - {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] [measurable_space E] [borel_space E] - [finite_dimensional ℝ E] (μ : measure E) [is_add_haar_measure μ] (f : E ≃ₗ[ℝ] E) (s : set E) : μ (f ⁻¹' s) = ennreal.of_real (abs (f.symm : E →ₗ[ℝ] E).det) * μ s := begin @@ -272,8 +270,6 @@ end /-- The preimage of a set `s` under a continuous linear equiv `f` has measure equal to `μ s` times the absolute value of the inverse of the determinant of `f`. -/ @[simp] lemma add_haar_preimage_continuous_linear_equiv - {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] [measurable_space E] [borel_space E] - [finite_dimensional ℝ E] (μ : measure E) [is_add_haar_measure μ] (f : E ≃L[ℝ] E) (s : set E) : μ (f ⁻¹' s) = ennreal.of_real (abs (f.symm : E →ₗ[ℝ] E).det) * μ s := add_haar_preimage_linear_equiv μ _ s @@ -281,8 +277,6 @@ add_haar_preimage_linear_equiv μ _ s /-- The image of a set `s` under a linear map `f` has measure equal to `μ s` times the absolute value of the determinant of `f`. -/ @[simp] lemma add_haar_image_linear_map - {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] [measurable_space E] [borel_space E] - [finite_dimensional ℝ E] (μ : measure E) [is_add_haar_measure μ] (f : E →ₗ[ℝ] E) (s : set E) : μ (f '' s) = ennreal.of_real (abs f.det) * μ s := begin @@ -303,8 +297,6 @@ end /-- The image of a set `s` under a continuous linear map `f` has measure equal to `μ s` times the absolute value of the determinant of `f`. -/ @[simp] lemma add_haar_image_continuous_linear_map - {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] [measurable_space E] [borel_space E] - [finite_dimensional ℝ E] (μ : measure E) [is_add_haar_measure μ] (f : E →L[ℝ] E) (s : set E) : μ (f '' s) = ennreal.of_real (abs (f : E →ₗ[ℝ] E).det) * μ s := add_haar_image_linear_map μ _ s @@ -312,8 +304,6 @@ add_haar_image_linear_map μ _ s /-- The image of a set `s` under a continuous linear equiv `f` has measure equal to `μ s` times the absolute value of the determinant of `f`. -/ @[simp] lemma add_haar_image_continuous_linear_equiv - {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] [measurable_space E] [borel_space E] - [finite_dimensional ℝ E] (μ : measure E) [is_add_haar_measure μ] (f : E ≃L[ℝ] E) (s : set E) : μ (f '' s) = ennreal.of_real (abs (f : E →ₗ[ℝ] E).det) * μ s := μ.add_haar_image_linear_map (f : E →ₗ[ℝ] E) s @@ -322,10 +312,6 @@ equal to `μ s` times the absolute value of the determinant of `f`. -/ ### Basic properties of Haar measures on real vector spaces -/ -variables {E : Type*} [normed_add_comm_group E] [measurable_space E] [normed_space ℝ E] - [finite_dimensional ℝ E] [borel_space E] (μ : measure E) [is_add_haar_measure μ] - {F : Type*} [normed_add_comm_group F] [normed_space ℝ F] [complete_space F] - lemma map_add_haar_smul {r : ℝ} (hr : r ≠ 0) : measure.map ((•) r) μ = ennreal.of_real (abs (r ^ (finrank ℝ E))⁻¹) • μ := begin @@ -566,6 +552,56 @@ begin simp only [real.to_nnreal_bit0, real.to_nnreal_one, le_refl], end +section + +/-! +### The Lebesgue measure associated to an alternating map +-/ + +variables {ι G : Type*} [fintype ι] [decidable_eq ι] +[normed_add_comm_group G] [normed_space ℝ G] [measurable_space G] [borel_space G] + +lemma add_haar_parallelepiped (b : basis ι ℝ G) (v : ι → G) : + b.add_haar (parallelepiped v) = ennreal.of_real (|b.det v|) := +begin + haveI : finite_dimensional ℝ G, from finite_dimensional.of_fintype_basis b, + have A : parallelepiped v = (b.constr ℕ v) '' (parallelepiped b), + { rw image_parallelepiped, + congr' 1 with i, + exact (b.constr_basis ℕ v i).symm }, + rw [A, add_haar_image_linear_map, basis.add_haar_self, mul_one, + ← linear_map.det_to_matrix b, ← basis.to_matrix_eq_to_matrix_constr], + refl, +end + +variables [finite_dimensional ℝ G] {n : ℕ} [_i : fact (finrank ℝ G = n)] +include _i + +/-- The Lebesgue measure associated to an alternating map. It gives measure `|ω v|` to the +parallelepiped spanned by the vectors `v₁, ..., vₙ`. Note that it is not always a Haar measure, +as it can be zero, but it is always locally finite and translation invariant. -/ +@[irreducible] noncomputable def _root_.alternating_map.measure + (ω : alternating_map ℝ G ℝ (fin n)) : measure G := +‖ω (fin_basis_of_finrank_eq ℝ G _i.out)‖₊ • (fin_basis_of_finrank_eq ℝ G _i.out).add_haar + +lemma _root_.alternating_map.measure_parallelepiped + (ω : alternating_map ℝ G ℝ (fin n)) (v : fin n → G) : + ω.measure (parallelepiped v) = ennreal.of_real (|ω v|) := +begin + conv_rhs { rw ω.eq_smul_basis_det (fin_basis_of_finrank_eq ℝ G _i.out) }, + simp only [add_haar_parallelepiped, alternating_map.measure, coe_nnreal_smul_apply, + alternating_map.smul_apply, algebra.id.smul_eq_mul, abs_mul, + ennreal.of_real_mul (abs_nonneg _), real.ennnorm_eq_of_real_abs] +end + +instance (ω : alternating_map ℝ G ℝ (fin n)) : is_add_left_invariant ω.measure := +by { rw [alternating_map.measure], apply_instance } + +instance (ω : alternating_map ℝ G ℝ (fin n)) : is_locally_finite_measure ω.measure := +by { rw [alternating_map.measure], apply_instance } + +end + /-! ### Density points diff --git a/src/measure_theory/measure/haar_of_basis.lean b/src/measure_theory/measure/haar_of_basis.lean new file mode 100644 index 0000000000000..e8da34d00525d --- /dev/null +++ b/src/measure_theory/measure/haar_of_basis.lean @@ -0,0 +1,160 @@ +/- +Copyright (c) 2022 Sébastien Gouëzel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sébastien Gouëzel +-/ +import measure_theory.measure.haar +import analysis.inner_product_space.pi_L2 + +/-! +# Additive Haar measure constructed from a basis + +Given a basis of a finite-dimensional real vector space, we define the corresponding Lebesgue +measure, which gives measure `1` to the parallelepiped spanned by the basis. + +## Main definitions + +* `parallelepiped v` is the parallelepiped spanned by a finite family of vectors. +* `basis.parallelepiped` is the parallelepiped associated to a basis, seen as a compact set with +nonempty interior. +* `basis.add_haar` is the Lebesgue measure associated to a basis, giving measure `1` to the +corresponding parallelepiped. + +In particular, we declare a `measure_space` instance on any finite-dimensional inner product space, +by using the Lebesgue measure associated to some orthonormal basis (which is in fact independent +of the basis). +-/ + +open set topological_space measure_theory measure_theory.measure finite_dimensional +open_locale big_operators + +noncomputable theory + +variables {ι ι' E F : Type*} [fintype ι] [fintype ι'] + +section add_comm_group + +variables [add_comm_group E] [module ℝ E] [add_comm_group F] [module ℝ F] + +/-- The closed parallelepiped spanned by a finite family of vectors. -/ +def parallelepiped (v : ι → E) : set E := +(λ (t : ι → ℝ), ∑ i, t i • v i) '' (Icc 0 1) + +lemma mem_parallelepiped_iff (v : ι → E) (x : E) : + x ∈ parallelepiped v ↔ ∃ (t : ι → ℝ) (ht : t ∈ Icc (0 : ι → ℝ) 1), x = ∑ i, t i • v i := +by simp [parallelepiped, eq_comm] + +lemma image_parallelepiped (f : E →ₗ[ℝ] F) (v : ι → E) : + f '' (parallelepiped v) = parallelepiped (f ∘ v) := +begin + simp only [parallelepiped, ← image_comp], + congr' 1 with t, + simp only [function.comp_app, linear_map.map_sum, linear_map.map_smulₛₗ, ring_hom.id_apply], +end + +/-- Reindexing a family of vectors does not change their parallelepiped. -/ +@[simp] lemma parallelepiped_comp_equiv (v : ι → E) (e : ι' ≃ ι) : + parallelepiped (v ∘ e) = parallelepiped v := +begin + simp only [parallelepiped], + let K : (ι' → ℝ) ≃ (ι → ℝ) := equiv.Pi_congr_left' (λ (a : ι'), ℝ) e, + have : Icc (0 : (ι → ℝ)) 1 = K '' (Icc (0 : (ι' → ℝ)) 1), + { rw ← equiv.preimage_eq_iff_eq_image, + ext x, + simp only [mem_preimage, mem_Icc, pi.le_def, pi.zero_apply, equiv.Pi_congr_left'_apply, + pi.one_apply], + refine ⟨λ h, ⟨λ i, _, λ i, _⟩, λ h, ⟨λ i, h.1 (e.symm i), λ i, h.2 (e.symm i)⟩⟩, + { simpa only [equiv.symm_apply_apply] using h.1 (e i) }, + { simpa only [equiv.symm_apply_apply] using h.2 (e i) } }, + rw [this, ← image_comp], + congr' 1 with x, + simpa only [orthonormal_basis.coe_reindex, function.comp_app, equiv.symm_apply_apply, + equiv.Pi_congr_left'_apply, equiv.apply_symm_apply] + using (e.symm.sum_comp (λ (i : ι'), x i • v (e i))).symm, +end + +/- The parallelepiped associated to an orthonormal basis of `ℝ` is either `[0, 1]` or `[-1, 0]`. -/ +lemma parallelepiped_orthonormal_basis_one_dim (b : orthonormal_basis ι ℝ ℝ) : + parallelepiped b = Icc 0 1 ∨ parallelepiped b = Icc (-1) 0 := +begin + have e : ι ≃ fin 1, + { apply fintype.equiv_fin_of_card_eq, + simp only [← finrank_eq_card_basis b.to_basis, finrank_self] }, + have B : parallelepiped (b.reindex e) = parallelepiped b, + { convert parallelepiped_comp_equiv b e.symm, + ext i, + simp only [orthonormal_basis.coe_reindex] }, + rw ← B, + let F : ℝ → (fin 1 → ℝ) := λ t, (λ i, t), + have A : Icc (0 : fin 1 → ℝ) 1 = F '' (Icc (0 : ℝ) 1), + { apply subset.antisymm, + { assume x hx, + refine ⟨x 0, ⟨hx.1 0, hx.2 0⟩, _⟩, + ext j, + simp only [subsingleton.elim j 0] }, + { rintros x ⟨y, hy, rfl⟩, + exact ⟨λ j, hy.1, λ j, hy.2⟩ } }, + rcases orthonormal_basis_one_dim (b.reindex e) with H|H, + { left, + simp only [H, parallelepiped, algebra.id.smul_eq_mul, mul_one, A, + finset.sum_singleton, ←image_comp, image_id', finset.univ_unique], }, + { right, + simp only [H, parallelepiped, algebra.id.smul_eq_mul, mul_one], + rw A, + simp only [←image_comp, mul_neg, mul_one, finset.sum_singleton, image_neg, preimage_neg_Icc, + neg_zero, finset.univ_unique] }, +end + +end add_comm_group + +section normed_space + +variables [normed_add_comm_group E] [normed_space ℝ E] + +/-- The parallelepiped spanned by a basis, as a compact set with nonempty interior. -/ +def basis.parallelepiped (b : basis ι ℝ E) : positive_compacts E := +{ carrier := parallelepiped b, + is_compact' := is_compact_Icc.image (continuous_finset_sum finset.univ + (λ (i : ι) (H : i ∈ finset.univ), (continuous_apply i).smul continuous_const)), + interior_nonempty' := + begin + suffices H : set.nonempty (interior (b.equiv_funL.symm.to_homeomorph '' (Icc 0 1))), + { dsimp only [parallelepiped], + convert H, + ext t, + exact (b.equiv_fun_symm_apply t).symm }, + have A : set.nonempty (interior (Icc (0 : ι → ℝ) 1)), + { rw [← pi_univ_Icc, interior_pi_set (@finite_univ ι _)], + simp only [univ_pi_nonempty_iff, pi.zero_apply, pi.one_apply, interior_Icc, nonempty_Ioo, + zero_lt_one, implies_true_iff] }, + rwa [← homeomorph.image_interior, nonempty_image_iff], + end } + +variables [measurable_space E] [borel_space E] + +/-- The Lebesgue measure associated to a basis, giving measure `1` to the parallelepiped spanned +by the basis. -/ +@[irreducible] def basis.add_haar (b : basis ι ℝ E) : measure E := +measure.add_haar_measure b.parallelepiped + +instance is_add_haar_measure_basis_add_haar (b : basis ι ℝ E) : + is_add_haar_measure b.add_haar := +by { rw basis.add_haar, exact measure.is_add_haar_measure_add_haar_measure _ } + +lemma basis.add_haar_self (b : basis ι ℝ E) : b.add_haar (parallelepiped b) = 1 := +by { rw [basis.add_haar], exact add_haar_measure_self } + +end normed_space + +/-- A finite dimensional inner product space has a canonical measure, the Lebesgue measure giving +volume `1` to the parallelepiped spanned by any orthonormal basis. We define the measure using +some arbitrary choice of orthonormal basis. The fact that it works with any orthonormal basis +is proved in `orthonormal_basis.volume_parallelepiped`. -/ +@[priority 100] instance measure_space_of_inner_product_space + [inner_product_space ℝ E] [finite_dimensional ℝ E] [measurable_space E] [borel_space E] : + measure_space E := +{ volume := (std_orthonormal_basis ℝ E).to_basis.add_haar } + +/- This instance should not be necessary, but Lean has difficulties to find it in product +situations if we do not declare it explicitly. -/ +instance real.measure_space : measure_space ℝ := by apply_instance diff --git a/src/measure_theory/measure/haar_of_inner.lean b/src/measure_theory/measure/haar_of_inner.lean new file mode 100644 index 0000000000000..b10f73db4734b --- /dev/null +++ b/src/measure_theory/measure/haar_of_inner.lean @@ -0,0 +1,69 @@ +/- +Copyright (c) 2022 Sébastien Gouëzel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sébastien Gouëzel +-/ +import analysis.inner_product_space.orientation +import measure_theory.measure.haar_lebesgue + +/-! +# Volume forms and measures on inner product spaces + +A volume form induces a Lebesgue measure on general finite-dimensional real vector spaces. In this +file, we discuss the specific situation of inner product spaces, where an orientation gives +rise to a canonical volume form. We show that the measure coming from this volume form gives +measure `1` to the parallelepiped spanned by any orthonormal basis, and that it coincides with +the canonical `volume` from the `measure_space` instance. +-/ + +open finite_dimensional measure_theory measure_theory.measure set + +variables {ι F : Type*} [fintype ι] [inner_product_space ℝ F] [finite_dimensional ℝ F] +[measurable_space F] [borel_space F] + +section + +variables {m n : ℕ} [_i : fact (finrank ℝ F = n)] +include _i + +/-- The volume form coming from an orientation in an inner product space gives measure `1` to the +parallelepiped associated to any orthonormal basis. This is a rephrasing of +`abs_volume_form_apply_of_orthonormal` in terms of measures. -/ +lemma orientation.measure_orthonormal_basis + (o : orientation ℝ F (fin n)) (b : orthonormal_basis ι ℝ F) : + o.volume_form.measure (parallelepiped b) = 1 := +begin + have e : ι ≃ fin n, + { refine fintype.equiv_fin_of_card_eq _, + rw [← _i.out, finrank_eq_card_basis b.to_basis] }, + have A : ⇑b = (b.reindex e) ∘ e, + { ext x, + simp only [orthonormal_basis.coe_reindex, function.comp_app, equiv.symm_apply_apply] }, + rw [A, parallelepiped_comp_equiv, alternating_map.measure_parallelepiped, + o.abs_volume_form_apply_of_orthonormal, ennreal.of_real_one], +end + +/-- In an oriented inner product space, the measure coming from the canonical volume form +associated to an orientation coincides with the volume. -/ +lemma orientation.measure_eq_volume (o : orientation ℝ F (fin n)) : + o.volume_form.measure = volume := +begin + have A : o.volume_form.measure ((std_orthonormal_basis ℝ F).to_basis.parallelepiped) = 1, + from orientation.measure_orthonormal_basis o (std_orthonormal_basis ℝ F), + rw [add_haar_measure_unique o.volume_form.measure + ((std_orthonormal_basis ℝ F).to_basis.parallelepiped), A, one_smul], + simp only [volume, basis.add_haar], +end + +end + +/-- The volume measure in a finite-dimensional inner product space gives measure `1` to the +parallelepiped spanned by any orthonormal basis. -/ +lemma orthonormal_basis.volume_parallelepiped (b : orthonormal_basis ι ℝ F) : + volume (parallelepiped b) = 1 := +begin + haveI : fact (finrank ℝ F = finrank ℝ F) := ⟨rfl⟩, + let o := (std_orthonormal_basis ℝ F).to_basis.orientation, + rw ← o.measure_eq_volume, + exact o.measure_orthonormal_basis b, +end diff --git a/src/measure_theory/measure/lebesgue.lean b/src/measure_theory/measure/lebesgue.lean index 71aca319c9d00..96f968b6af59a 100644 --- a/src/measure_theory/measure/lebesgue.lean +++ b/src/measure_theory/measure/lebesgue.lean @@ -9,13 +9,15 @@ import linear_algebra.matrix.diagonal import linear_algebra.matrix.transvection import measure_theory.constructions.pi import measure_theory.measure.stieltjes +import measure_theory.measure.haar_of_basis /-! # Lebesgue measure on the real line and on `ℝⁿ` -We construct Lebesgue measure on the real line, as a particular case of Stieltjes measure associated -to the function `x ↦ x`. We obtain as a consequence Lebesgue measure on `ℝⁿ`. We prove that they -are translation invariant. +We show that the Lebesgue measure on the real line (constructed as a particular case of additive +Haar measure on inner product spaces) coincides with the Stieltjes measure associated +to the function `x ↦ x`. We deduce properties of this measure on `ℝ`, and then of the product +Lebesgue measure on `ℝⁿ`. In particular, we prove that they are translation invariant. We show that, on `ℝⁿ`, a linear map acts on Lebesgue measure by rescaling it through the absolute value of its determinant, in `real.map_linear_map_volume_pi_eq_smul_volume_pi`. @@ -33,17 +35,31 @@ open_locale big_operators ennreal nnreal topological_space ### Definition of the Lebesgue measure and lengths of intervals -/ -/-- Lebesgue measure on the Borel sigma algebra, giving measure `b - a` to the interval `[a, b]`. -/ -instance real.measure_space : measure_space ℝ := -⟨stieltjes_function.id.measure⟩ - namespace real variables {ι : Type*} [fintype ι] -open_locale topological_space +/-- The volume on the real line (as a particular case of the volume on a finite-dimensional +inner product space) coincides with the Stieltjes measure coming from the identity function. -/ +lemma volume_eq_stieltjes_id : (volume : measure ℝ) = stieltjes_function.id.measure := +begin + haveI : is_add_left_invariant stieltjes_function.id.measure := + ⟨λ a, eq.symm $ real.measure_ext_Ioo_rat $ λ p q, + by simp only [measure.map_apply (measurable_const_add a) measurable_set_Ioo, + sub_sub_sub_cancel_right, stieltjes_function.measure_Ioo, stieltjes_function.id_left_lim, + stieltjes_function.id_apply, id.def, preimage_const_add_Ioo]⟩, + have A : stieltjes_function.id.measure (std_orthonormal_basis ℝ ℝ).to_basis.parallelepiped = 1, + { change stieltjes_function.id.measure (parallelepiped (std_orthonormal_basis ℝ ℝ)) = 1, + rcases parallelepiped_orthonormal_basis_one_dim (std_orthonormal_basis ℝ ℝ) with H|H; + simp only [H, stieltjes_function.measure_Icc, stieltjes_function.id_apply, id.def, tsub_zero, + stieltjes_function.id_left_lim, sub_neg_eq_add, zero_add, ennreal.of_real_one] }, + conv_rhs { rw [add_haar_measure_unique stieltjes_function.id.measure + (std_orthonormal_basis ℝ ℝ).to_basis.parallelepiped, A] }, + simp only [volume, basis.add_haar, one_smul], +end -theorem volume_val (s) : volume s = stieltjes_function.id.measure s := rfl +theorem volume_val (s) : volume s = stieltjes_function.id.measure s := +by simp [volume_eq_stieltjes_id] @[simp] lemma volume_Ico {a b : ℝ} : volume (Ico a b) = of_real (b - a) := by simp [volume_val] @@ -132,6 +148,23 @@ instance is_finite_measure_restrict_Ioc (x y : ℝ) : is_finite_measure (volume. instance is_finite_measure_restrict_Ioo (x y : ℝ) : is_finite_measure (volume.restrict (Ioo x y)) := ⟨by simp⟩ +lemma volume_le_diam (s : set ℝ) : volume s ≤ emetric.diam s := +begin + by_cases hs : metric.bounded s, + { rw [real.ediam_eq hs, ← volume_Icc], + exact volume.mono (real.subset_Icc_Inf_Sup_of_bounded hs) }, + { rw metric.ediam_of_unbounded hs, exact le_top } +end + +lemma _root_.filter.eventually.volume_pos_of_nhds_real + {p : ℝ → Prop} {a : ℝ} (h : ∀ᶠ x in 𝓝 a, p x) : + (0 : ℝ≥0∞) < volume {x | p x} := +begin + rcases h.exists_Ioo_subset with ⟨l, u, hx, hs⟩, + refine lt_of_lt_of_le _ (measure_mono hs), + simpa [-mem_Ioo] using hx.1.trans hx.2 +end + /-! ### Volume of a box in `ℝⁿ` -/ @@ -184,14 +217,6 @@ begin exact (ennreal.of_real_pow (mul_nonneg zero_le_two hr) _).symm end -lemma volume_le_diam (s : set ℝ) : volume s ≤ emetric.diam s := -begin - by_cases hs : metric.bounded s, - { rw [real.ediam_eq hs, ← volume_Icc], - exact volume.mono (real.subset_Icc_Inf_Sup_of_bounded hs) }, - { rw metric.ediam_of_unbounded hs, exact le_top } -end - lemma volume_pi_le_prod_diam (s : set (ι → ℝ)) : volume s ≤ ∏ i : ι, emetric.diam (function.eval i '' s) := calc volume s ≤ volume (pi univ (λ i, closure (function.eval i '' s))) : @@ -210,14 +235,9 @@ calc volume s ≤ ∏ i : ι, emetric.diam (function.eval i '' s) : volume_pi_le by simp only [ennreal.coe_one, one_mul, finset.prod_const, fintype.card] /-! -### Images of the Lebesgue measure under translation/multiplication in ℝ +### Images of the Lebesgue measure under multiplication in ℝ -/ -instance is_add_left_invariant_real_volume : - is_add_left_invariant (volume : measure ℝ) := -⟨λ a, eq.symm $ real.measure_ext_Ioo_rat $ λ p q, - by simp [measure.map_apply (measurable_const_add a) measurable_set_Ioo, sub_sub_sub_cancel_right]⟩ - lemma smul_map_volume_mul_left {a : ℝ} (h : a ≠ 0) : ennreal.of_real (|a|) • measure.map ((*) a) volume = volume := begin @@ -258,10 +278,6 @@ calc volume ((* a) ⁻¹' s) = measure.map (* a) volume s : ((homeomorph.mul_right₀ a h).to_measurable_equiv.map_apply s).symm ... = ennreal.of_real (abs a⁻¹) * volume s : by { rw map_volume_mul_right h, refl } -instance : is_neg_invariant (volume : measure ℝ) := -⟨eq.symm $ real.measure_ext_Ioo_rat $ λ p q, by simp [show volume.neg (Ioo (p : ℝ) q) = _, - from measure.map_apply measurable_neg measurable_set_Ioo]⟩ - /-! ### Images of the Lebesgue measure under translation/linear maps in ℝⁿ -/ @@ -376,20 +392,8 @@ end end real -open_locale topological_space - -lemma filter.eventually.volume_pos_of_nhds_real {p : ℝ → Prop} {a : ℝ} (h : ∀ᶠ x in 𝓝 a, p x) : - (0 : ℝ≥0∞) < volume {x | p x} := -begin - rcases h.exists_Ioo_subset with ⟨l, u, hx, hs⟩, - refine lt_of_lt_of_le _ (measure_mono hs), - simpa [-mem_Ioo] using hx.1.trans hx.2 -end - section region_between -open_locale classical - variable {α : Type*} /-- The region between two real-valued functions on an arbitrary set. -/ @@ -457,6 +461,7 @@ theorem volume_region_between_eq_lintegral' (hf : measurable f) (hg : measurable g) (hs : measurable_set s) : μ.prod volume (region_between f g s) = ∫⁻ y in s, ennreal.of_real ((g - f) y) ∂μ := begin + classical, rw measure.prod_apply, { have h : (λ x, volume {a | x ∈ s ∧ a ∈ Ioo (f x) (g x)}) = s.indicator (λ x, ennreal.of_real (g x - f x)), diff --git a/src/measure_theory/measure/measure_space.lean b/src/measure_theory/measure/measure_space.lean index ef343796071df..82ab393b11391 100644 --- a/src/measure_theory/measure/measure_space.lean +++ b/src/measure_theory/measure/measure_space.lean @@ -1022,6 +1022,10 @@ begin simp [map_of_not_ae_measurable hf, map_of_not_ae_measurable hfc] } end +@[simp] protected lemma map_smul_nnreal (c : ℝ≥0) (μ : measure α) (f : α → β) : + (c • μ).map f = c • μ.map f := +μ.map_smul (c : ℝ≥0∞) f + /-- We can evaluate the pushforward on measurable sets. For non-measurable sets, see `measure_theory.measure.le_map_apply` and `measurable_equiv.map_apply`. -/ @[simp] theorem map_apply_of_ae_measurable From ee0c179cd3c8a45aa5bffbf1b41d8dbede452865 Mon Sep 17 00:00:00 2001 From: Riccardo Brasca Date: Tue, 13 Dec 2022 15:36:28 +0000 Subject: [PATCH 0030/1291] feat(number_theory/number_field/norm): add file (#17672) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We add `norm' : (𝓞 L) →* (𝓞 K)`, that is `algebra.norm K` as a morphism between the rings of integers. From flt-regular --- src/number_theory/number_field/norm.lean | 71 ++++++++++++++++++++++++ src/ring_theory/norm.lean | 2 +- 2 files changed, 72 insertions(+), 1 deletion(-) create mode 100644 src/number_theory/number_field/norm.lean diff --git a/src/number_theory/number_field/norm.lean b/src/number_theory/number_field/norm.lean new file mode 100644 index 0000000000000..35428dcf00f67 --- /dev/null +++ b/src/number_theory/number_field/norm.lean @@ -0,0 +1,71 @@ +/- +Copyright (c) 2022 Riccardo Brasca. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Riccardo Brasca, Eric Rodriguez +-/ + +import number_theory.number_field.basic +import ring_theory.norm + +/-! +# Norm in number fields +Given a finite extension of number fields, we define the norm morphism as a function between the +rings of integers. + +## Main definitions +* `ring_of_integers.norm K` : `algebra.norm` as a morphism `(𝓞 L) →* (𝓞 K)`. +## Main results +* `algebra.dvd_norm` : if `L/K` is a finite Galois extension of fields, then, for all `(x : 𝓞 L)` + we have that `x ∣ algebra_map (𝓞 K) (𝓞 L) (norm K x)`. + +-/ + +open_locale number_field big_operators + +open finset number_field algebra + +namespace ring_of_integers + +variables {L : Type*} (K : Type*) [field K] [field L] [algebra K L] [finite_dimensional K L] + +/-- `algebra.norm` as a morphism betwen the rings of integers. -/ +@[simps] noncomputable def norm [is_separable K L] : (𝓞 L) →* (𝓞 K) := +((algebra.norm K).restrict (𝓞 L)).cod_restrict (𝓞 K) (λ x, is_integral_norm K x.2) + +local attribute [instance] number_field.ring_of_integers_algebra + +lemma coe_algebra_map_norm [is_separable K L] (x : 𝓞 L) : + (algebra_map (𝓞 K) (𝓞 L) (norm K x) : L) = algebra_map K L (algebra.norm K (x : L)) := rfl + +lemma is_unit_norm [is_galois K L] {x : 𝓞 L} : + is_unit (norm K x) ↔ is_unit x := +begin + classical, + refine ⟨λ hx, _, is_unit.map _⟩, + replace hx : is_unit (algebra_map (𝓞 K) (𝓞 L) $ norm K x) := hx.map (algebra_map (𝓞 K) $ 𝓞 L), + refine @is_unit_of_mul_is_unit_right (𝓞 L) _ + ⟨(univ \ { alg_equiv.refl }).prod (λ (σ : L ≃ₐ[K] L), σ x), + prod_mem (λ σ hσ, map_is_integral (σ : L →+* L).to_int_alg_hom x.2)⟩ _ _, + convert hx using 1, + ext, + push_cast, + convert_to (univ \ { alg_equiv.refl }).prod (λ (σ : L ≃ₐ[K] L), σ x) * (∏ (σ : L ≃ₐ[K] L) in + {alg_equiv.refl}, σ (x : L)) = _, + { rw [prod_singleton, alg_equiv.coe_refl, id] }, + { rw [prod_sdiff $ subset_univ _, ←norm_eq_prod_automorphisms, coe_algebra_map_norm] } +end + +/-- If `L/K` is a finite Galois extension of fields, then, for all `(x : 𝓞 L)` we have that +`x ∣ algebra_map (𝓞 K) (𝓞 L) (norm K x)`. -/ +lemma dvd_norm [is_galois K L] (x : 𝓞 L) : x ∣ algebra_map (𝓞 K) (𝓞 L) (norm K x) := +begin + classical, + have hint : (∏ (σ : L ≃ₐ[K] L) in univ.erase alg_equiv.refl, σ x) ∈ 𝓞 L := + subalgebra.prod_mem _ (λ σ hσ, (mem_ring_of_integers _ _).2 + (map_is_integral σ (ring_of_integers.is_integral_coe x))), + refine ⟨⟨_, hint⟩, subtype.ext _⟩, + rw [coe_algebra_map_norm K x, norm_eq_prod_automorphisms], + simp [← finset.mul_prod_erase _ _ (mem_univ alg_equiv.refl)] +end + +end ring_of_integers diff --git a/src/ring_theory/norm.lean b/src/ring_theory/norm.lean index a9ee578b28db7..9cd10eb83f041 100644 --- a/src/ring_theory/norm.lean +++ b/src/ring_theory/norm.lean @@ -287,7 +287,7 @@ begin exact is_separable.separable K _ } end -lemma norm_eq_prod_automorphisms [finite_dimensional K L] [is_galois K L] {x : L}: +lemma norm_eq_prod_automorphisms [finite_dimensional K L] [is_galois K L] (x : L) : algebra_map K L (norm K x) = ∏ (σ : L ≃ₐ[K] L), σ x := begin apply no_zero_smul_divisors.algebra_map_injective L (algebraic_closure L), From 41a3880274a384b8508802566a10c931e039e55d Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 13 Dec 2022 19:33:24 +0000 Subject: [PATCH 0031/1291] chore(data/set/sigma): Fix lemma (#17929) I accidentally used a non-bound variable where I wanted to bind it. This resulted in an unusable lemma. --- src/data/set/sigma.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/data/set/sigma.lean b/src/data/set/sigma.lean index 8355fb6a3bb7e..1aa3bfb920dd8 100644 --- a/src/data/set/sigma.lean +++ b/src/data/set/sigma.lean @@ -67,7 +67,7 @@ lemma exists_sigma_iff {p : (Σ i, α i) → Prop} : (∃ x ∈ s.sigma t, p x) ↔ ∃ (i ∈ s) (a ∈ t i), p ⟨i, a⟩ := ⟨λ ⟨⟨i, a⟩, ha, h⟩, ⟨i, ha.1, a, ha.2, h⟩, λ ⟨i, hi, a, ha, h⟩, ⟨⟨i, a⟩, ⟨hi, ha⟩, h⟩⟩ -@[simp] lemma sigma_empty : s.sigma (λ _, (∅ : set (α i))) = ∅ := ext $ λ _, and_false _ +@[simp] lemma sigma_empty : s.sigma (λ i, (∅ : set (α i))) = ∅ := ext $ λ _, and_false _ @[simp] lemma empty_sigma : (∅ : set ι).sigma t = ∅ := ext $ λ _, false_and _ lemma univ_sigma_univ : (@univ ι).sigma (λ _, @univ (α i)) = univ := ext $ λ _, true_and _ @[simp] lemma sigma_univ : s.sigma (λ _, univ : Π i, set (α i)) = sigma.fst ⁻¹' s := From 4a1bf94891f08983bb49f7ac53b85ecf8c90250a Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 13 Dec 2022 22:24:40 +0000 Subject: [PATCH 0032/1291] =?UTF-8?q?feat(data/set/n=5Fary):=20Distributiv?= =?UTF-8?q?ity=20of=20`=E2=88=A9`=20(#17924)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit `set.image2 f` for `injective2 f` distributes over intersection. Also move the required results from `data.set.prod` to `data.set.n_ary`. As a bonus, this makes quite a few files not depend on `data.set.n_ary` anymore and matches the import direction for the corresponding `finset` files. --- src/data/finset/n_ary.lean | 8 ++++++++ src/data/set/n_ary.lean | 27 +++++++++++++++++++++++++-- src/data/set/prod.lean | 17 +---------------- src/order/bounds/basic.lean | 1 + 4 files changed, 35 insertions(+), 18 deletions(-) diff --git a/src/data/finset/n_ary.lean b/src/data/finset/n_ary.lean index f960ff37feea1..2ce67aac5c1c9 100644 --- a/src/data/finset/n_ary.lean +++ b/src/data/finset/n_ary.lean @@ -109,6 +109,14 @@ 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_left [decidable_eq α] (hf : injective2 f) : + image₂ f (s ∩ s') t = image₂ f s t ∩ image₂ f s' t := +coe_injective $ by { push_cast, exact image2_inter_left hf } + +lemma image₂_inter_right [decidable_eq β] (hf : injective2 f) : + image₂ f s (t ∩ t') = image₂ f s t ∩ image₂ f s t' := +coe_injective $ by { push_cast, exact image2_inter_right hf } + 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 } diff --git a/src/data/set/n_ary.lean b/src/data/set/n_ary.lean index 7ec9424a28b02..5c26c45a1cefd 100644 --- a/src/data/set/n_ary.lean +++ b/src/data/set/n_ary.lean @@ -3,8 +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 -/ -import data.set.basic -import data.set.image +import data.set.prod /-! # N-ary images of sets @@ -64,6 +63,24 @@ lemma forall_image2_iff {p : γ → Prop} : image2 f s t ⊆ u ↔ ∀ (x ∈ s) (y ∈ t), f x y ∈ u := forall_image2_iff +variables (f) + +@[simp] lemma image_prod : (λ x : α × β, f x.1 x.2) '' s ×ˢ t = image2 f s t := +ext $ λ a, +⟨ by { rintro ⟨_, _, rfl⟩, exact ⟨_, _, (mem_prod.1 ‹_›).1, (mem_prod.1 ‹_›).2, rfl⟩ }, + by { rintro ⟨_, _, _, _, rfl⟩, exact ⟨(_, _), ⟨‹_›, ‹_›⟩, rfl⟩ }⟩ + +@[simp] lemma image_uncurry_prod (s : set α) (t : set β) : uncurry f '' s ×ˢ t = image2 f s t := +image_prod _ + +@[simp] lemma image2_mk_eq_prod : image2 prod.mk s t = s ×ˢ t := ext $ by simp + +@[simp] lemma image2_curry (f : α × β → γ) (s : set α) (t : set β) : + image2 (λ a b, f (a, b)) s t = f '' s ×ˢ t := +by simp [←image_uncurry_prod, uncurry] + +variables {f} + lemma image2_union_left : image2 f (s ∪ s') t = image2 f s t ∪ image2 f s' t := begin ext c, @@ -82,6 +99,12 @@ begin simp [mem_union, *] } end +lemma image2_inter_left (hf : injective2 f) : image2 f (s ∩ s') t = image2 f s t ∩ image2 f s' t := +by simp_rw [←image_uncurry_prod, inter_prod, ←image_inter hf.uncurry] + +lemma image2_inter_right (hf : injective2 f) : image2 f s (t ∩ t') = image2 f s t ∩ image2 f s t' := +by simp_rw [←image_uncurry_prod, prod_inter, ←image_inter hf.uncurry] + @[simp] lemma image2_empty_left : image2 f ∅ t = ∅ := ext $ by simp @[simp] lemma image2_empty_right : image2 f s ∅ = ∅ := ext $ by simp diff --git a/src/data/set/prod.lean b/src/data/set/prod.lean index 73ba2701c9978..37c51a4e8ea71 100644 --- a/src/data/set/prod.lean +++ b/src/data/set/prod.lean @@ -3,8 +3,7 @@ Copyright (c) 2017 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Johannes Hölzl, Patrick Massot -/ -import data.set.basic -import data.set.n_ary +import data.set.image /-! # Sets in product and pi types @@ -291,20 +290,6 @@ begin refl, end -@[simp] lemma image_prod (f : α → β → γ) : (λ x : α × β, f x.1 x.2) '' s ×ˢ t = image2 f s t := -set.ext $ λ a, -⟨ by { rintro ⟨_, _, rfl⟩, exact ⟨_, _, (mem_prod.mp ‹_›).1, (mem_prod.mp ‹_›).2, rfl⟩ }, - by { rintro ⟨_, _, _, _, rfl⟩, exact ⟨(_, _), mem_prod.mpr ⟨‹_›, ‹_›⟩, rfl⟩ }⟩ - -@[simp] lemma image2_mk_eq_prod : image2 prod.mk s t = s ×ˢ t := ext $ by simp - -@[simp] lemma image2_curry (f : α × β → γ) (s : set α) (t : set β) : - image2 (λ a b, f (a, b)) s t = (s ×ˢ t).image f := -by rw [←image2_mk_eq_prod, image_image2] - -@[simp] lemma image_uncurry_prod (f : α → β → γ) (s : set α) (t : set β) : - uncurry f '' s ×ˢ t = image2 f s t := by { rw ←image2_curry, refl } - section mono variables [preorder α] {f : α → set β} {g : α → set γ} diff --git a/src/order/bounds/basic.lean b/src/order/bounds/basic.lean index 57ca0813a08fd..7f6acfdbd8296 100644 --- a/src/order/bounds/basic.lean +++ b/src/order/bounds/basic.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Yury Kudryashov -/ import data.set.intervals.basic +import data.set.n_ary /-! From 46ae64dc15f6dc8eda1654dd8e7faa3b21bb6cbe Mon Sep 17 00:00:00 2001 From: Adam Topaz Date: Tue, 13 Dec 2022 22:24:41 +0000 Subject: [PATCH 0033/1291] chore(group_theory/group_action/prod): Fix typo in `to_additive` attribute. (#17930) --- src/group_theory/group_action/prod.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/group_theory/group_action/prod.lean b/src/group_theory/group_action/prod.lean index 22232fa2d0a0f..728f6886acc79 100644 --- a/src/group_theory/group_action/prod.lean +++ b/src/group_theory/group_action/prod.lean @@ -52,7 +52,7 @@ by rw [prod.smul_mk, smul_zero] variables [has_pow α E] [has_pow β E] @[to_additive has_smul] instance has_pow : has_pow (α × β) E := { pow := λ p c, (p.1 ^ c, p.2 ^ c) } -@[simp, to_additive smul_snd, to_additive_reorder 6] +@[simp, to_additive smul_fst, to_additive_reorder 6] lemma pow_fst (p : α × β) (c : E) : (p ^ c).fst = p.fst ^ c := rfl @[simp, to_additive smul_snd, to_additive_reorder 6] lemma pow_snd (p : α × β) (c : E) : (p ^ c).snd = p.snd ^ c := rfl From f1adf6ad90b09ebecc1d9ce1a11de81a839a4515 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Tue, 13 Dec 2022 22:30:39 -0600 Subject: [PATCH 0034/1291] feat(ring_theory/power_series): `X s` commutes with any power series Use this fact to golf a proof and generalize `order_eq_multiplicity_X` from `[comm_semiring R]` to `[semiring R]`. Should we redefine `order` to be `multiplicity` in all cases? --- src/ring_theory/power_series/basic.lean | 26 ++++++++++++++++++------- 1 file changed, 19 insertions(+), 7 deletions(-) diff --git a/src/ring_theory/power_series/basic.lean b/src/ring_theory/power_series/basic.lean index e512ad9a7cc90..5752dc54e58b4 100644 --- a/src/ring_theory/power_series/basic.lean +++ b/src/ring_theory/power_series/basic.lean @@ -209,6 +209,16 @@ begin exact le_add_left le_rfl end +@[simp] lemma commute_monomial {a : R} {n} : + commute φ (monomial R n a) ↔ ∀ m, commute (coeff R m φ) a := +begin + refine ext_iff.trans ⟨λ h m, _, λ h m, _⟩, + { have := h (m + n), + rwa [coeff_add_mul_monomial, add_comm, coeff_add_monomial_mul] at this }, + { rw [coeff_mul_monomial, coeff_monomial_mul], + split_ifs; [apply h, refl] } +end + protected lemma one_mul : (1 : mv_power_series σ R) * φ = φ := ext $ λ n, by simpa using coeff_add_monomial_mul 0 n φ 1 @@ -323,6 +333,9 @@ coeff_monomial_same _ _ lemma coeff_zero_X (s : σ) : coeff R (0 : σ →₀ ℕ) (X s : mv_power_series σ R) = 0 := by { rw [coeff_X, if_neg], intro h, exact one_ne_zero (single_eq_zero.mp h.symm) } +lemma commute_X (φ : mv_power_series σ R) (s : σ) : commute φ (X s) := +φ.commute_monomial.mpr $ λ m, commute.one_right _ + lemma X_def (s : σ) : X s = monomial R (single s 1) 1 := rfl lemma X_pow_eq (s : σ) (n : ℕ) : @@ -353,11 +366,8 @@ begin end lemma coeff_zero_X_mul (φ : mv_power_series σ R) (s : σ) : - coeff R (0 : σ →₀ ℕ) (X s * φ) = 0 := -begin - have : ¬single s 1 ≤ 0, from λ h, by simpa using h s, - simp only [X, coeff_monomial_mul, if_neg this] -end + coeff R (0 : σ →₀ ℕ) (X s * φ) = 0 := +by rw [← (φ.commute_X s).eq, coeff_zero_mul_X] variables (σ) (R) @@ -1010,6 +1020,8 @@ variable {R} /-- The variable of the formal power series ring.-/ def X : power_series R := mv_power_series.X () +lemma commute_X (φ : power_series R) : commute φ X := φ.commute_X _ + @[simp] lemma coeff_zero_eq_constant_coeff : ⇑(coeff R 0) = constant_coeff R := by { rw [coeff, finsupp.single_zero], refl } @@ -1858,7 +1870,7 @@ begin simpa [part_enat.coe_lt_iff] using λ _, hn } end -lemma order_eq_multiplicity_X {R : Type*} [comm_semiring R] (φ : power_series R) : +lemma order_eq_multiplicity_X {R : Type*} [semiring R] (φ : power_series R) : order φ = multiplicity X φ := begin rcases eq_or_ne φ 0 with rfl|hφ, @@ -1872,7 +1884,7 @@ begin (order_finite_iff_ne_zero.mpr hφ)) (part_enat.find_le _ _ _), rintro ⟨ψ, H⟩, have := congr_arg (coeff R n) H, - rw [mul_comm, coeff_mul_of_lt_order, ←hn] at this, + rw [← (ψ.commute_X.pow_right _).eq, coeff_mul_of_lt_order, ←hn] at this, { exact coeff_order _ this }, { rw [X_pow_eq, order_monomial], split_ifs, From 27982536ee0a7748e5f20ac1646e94111a13480b Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Tue, 13 Dec 2022 22:32:36 -0600 Subject: [PATCH 0035/1291] Revert "feat(ring_theory/power_series): `X s` commutes with any power series" This reverts commit f1adf6ad90b09ebecc1d9ce1a11de81a839a4515. --- src/ring_theory/power_series/basic.lean | 26 +++++++------------------ 1 file changed, 7 insertions(+), 19 deletions(-) diff --git a/src/ring_theory/power_series/basic.lean b/src/ring_theory/power_series/basic.lean index 5752dc54e58b4..e512ad9a7cc90 100644 --- a/src/ring_theory/power_series/basic.lean +++ b/src/ring_theory/power_series/basic.lean @@ -209,16 +209,6 @@ begin exact le_add_left le_rfl end -@[simp] lemma commute_monomial {a : R} {n} : - commute φ (monomial R n a) ↔ ∀ m, commute (coeff R m φ) a := -begin - refine ext_iff.trans ⟨λ h m, _, λ h m, _⟩, - { have := h (m + n), - rwa [coeff_add_mul_monomial, add_comm, coeff_add_monomial_mul] at this }, - { rw [coeff_mul_monomial, coeff_monomial_mul], - split_ifs; [apply h, refl] } -end - protected lemma one_mul : (1 : mv_power_series σ R) * φ = φ := ext $ λ n, by simpa using coeff_add_monomial_mul 0 n φ 1 @@ -333,9 +323,6 @@ coeff_monomial_same _ _ lemma coeff_zero_X (s : σ) : coeff R (0 : σ →₀ ℕ) (X s : mv_power_series σ R) = 0 := by { rw [coeff_X, if_neg], intro h, exact one_ne_zero (single_eq_zero.mp h.symm) } -lemma commute_X (φ : mv_power_series σ R) (s : σ) : commute φ (X s) := -φ.commute_monomial.mpr $ λ m, commute.one_right _ - lemma X_def (s : σ) : X s = monomial R (single s 1) 1 := rfl lemma X_pow_eq (s : σ) (n : ℕ) : @@ -366,8 +353,11 @@ begin end lemma coeff_zero_X_mul (φ : mv_power_series σ R) (s : σ) : - coeff R (0 : σ →₀ ℕ) (X s * φ) = 0 := -by rw [← (φ.commute_X s).eq, coeff_zero_mul_X] + coeff R (0 : σ →₀ ℕ) (X s * φ) = 0 := +begin + have : ¬single s 1 ≤ 0, from λ h, by simpa using h s, + simp only [X, coeff_monomial_mul, if_neg this] +end variables (σ) (R) @@ -1020,8 +1010,6 @@ variable {R} /-- The variable of the formal power series ring.-/ def X : power_series R := mv_power_series.X () -lemma commute_X (φ : power_series R) : commute φ X := φ.commute_X _ - @[simp] lemma coeff_zero_eq_constant_coeff : ⇑(coeff R 0) = constant_coeff R := by { rw [coeff, finsupp.single_zero], refl } @@ -1870,7 +1858,7 @@ begin simpa [part_enat.coe_lt_iff] using λ _, hn } end -lemma order_eq_multiplicity_X {R : Type*} [semiring R] (φ : power_series R) : +lemma order_eq_multiplicity_X {R : Type*} [comm_semiring R] (φ : power_series R) : order φ = multiplicity X φ := begin rcases eq_or_ne φ 0 with rfl|hφ, @@ -1884,7 +1872,7 @@ begin (order_finite_iff_ne_zero.mpr hφ)) (part_enat.find_le _ _ _), rintro ⟨ψ, H⟩, have := congr_arg (coeff R n) H, - rw [← (ψ.commute_X.pow_right _).eq, coeff_mul_of_lt_order, ←hn] at this, + rw [mul_comm, coeff_mul_of_lt_order, ←hn] at this, { exact coeff_order _ this }, { rw [X_pow_eq, order_monomial], split_ifs, From b8858394c9190d253d98af50a8a06f3dda8f2d36 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 14 Dec 2022 05:27:13 +0000 Subject: [PATCH 0036/1291] refactor(data/set/intervals): Generalize `set.interval` to lattices (#17776) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Replace ``` def interval (a b : α) := Icc (min a b) (max a b) ``` by ``` def interval (a b : α) := Icc (a ⊓ b) (a ⊔ b) ``` Generalize lemmas accordingly. --- src/analysis/convex/quasiconvex.lean | 2 +- .../set/intervals/unordered_interval.lean | 141 ++++++++++-------- src/data/set/pointwise/interval.lean | 10 +- .../integral/interval_integral.lean | 2 +- src/measure_theory/measure/lebesgue.lean | 2 +- 5 files changed, 84 insertions(+), 73 deletions(-) diff --git a/src/analysis/convex/quasiconvex.lean b/src/analysis/convex/quasiconvex.lean index 78e8c890699a9..74c304828e03f 100644 --- a/src/analysis/convex/quasiconvex.lean +++ b/src/analysis/convex/quasiconvex.lean @@ -119,7 +119,7 @@ begin rw [quasilinear_on, quasiconvex_on_iff_le_max, quasiconcave_on_iff_min_le, and_and_and_comm, and_self], apply and_congr_right', - simp_rw [←forall_and_distrib, interval, mem_Icc, and_comm], + simp_rw [←forall_and_distrib, ←Icc_min_max, mem_Icc, and_comm], end lemma quasiconvex_on.convex_lt (hf : quasiconvex_on 𝕜 s f) (r : β) : convex 𝕜 {x ∈ s | f x < r} := diff --git a/src/data/set/intervals/unordered_interval.lean b/src/data/set/intervals/unordered_interval.lean index 61410b54b796e..f5fc402a3a656 100644 --- a/src/data/set/intervals/unordered_interval.lean +++ b/src/data/set/intervals/unordered_interval.lean @@ -9,14 +9,21 @@ import data.set.intervals.basic /-! # Intervals without endpoints ordering -In any decidable linear order `α`, we define the set of elements lying between two elements `a` and -`b` as `Icc (min a b) (max a b)`. +In any lattice `α`, we define `interval a b` to be `Icc (a ⊓ b) (a ⊔ b)`, which in a linear order is +the set of elements lying between `a` and `b`. `Icc a b` requires the assumption `a ≤ b` to be meaningful, which is sometimes inconvenient. The interval as defined in this file is always the set of things lying between `a` and `b`, regardless of the relative order of `a` and `b`. -For real numbers, `Icc (min a b) (max a b)` is the same as `segment ℝ a b`. +For real numbers, `interval a b` is the same as `segment ℝ a b`. + +In a product or pi type, `interval a b` is the smallest box containing `a` and `b`. For example, +`interval (1, -1) (-1, 1) = Icc (-1, -1) (1, 1)` is the square of vertices `(1, -1)`, `(-1, -1)`, +`(-1, 1)`, `(1, 1)`. + +In `finset α` (seen as a hypercube of dimension `fintype.card α`), `interval a b` is the smallest +subcube containing both `a` and `b`. ## Notation @@ -27,27 +34,28 @@ make the notation available. open function order_dual (to_dual of_dual) -namespace set +variables {α β : Type*} -section linear_order -variables {α β : Type*} [linear_order α] [linear_order β] {f : α → β} {s : set α} - {a a₁ a₂ b b₁ b₂ c x : α} +namespace set +section lattice +variables [lattice α] {a a₁ a₂ b b₁ b₂ c x : α} -/-- `interval a b` is the set of elements lying between `a` and `b`, with `a` and `b` included. -/ -def interval (a b : α) := Icc (min a b) (max a b) +/-- `interval a b` is the set of elements lying between `a` and `b`, with `a` and `b` included. +Note that we define it more generally in a lattice as `set.Icc (a ⊓ b) (a ⊔ b)`. In a product type, +`interval` corresponds to the bounding box of the two elements. -/ +def interval (a b : α) : set α := Icc (a ⊓ b) (a ⊔ b) localized "notation (name := set.interval) `[`a `, ` b `]` := set.interval a b" in interval @[simp] lemma dual_interval (a b : α) : [to_dual a, to_dual b] = of_dual ⁻¹' [a, b] := dual_Icc @[simp] lemma interval_of_le (h : a ≤ b) : [a, b] = Icc a b := -by rw [interval, min_eq_left h, max_eq_right h] +by rw [interval, inf_eq_left.2 h, sup_eq_right.2 h] @[simp] lemma interval_of_ge (h : b ≤ a) : [a, b] = Icc b a := -by { rw [interval, min_eq_right h, max_eq_left h] } +by rw [interval, inf_eq_right.2 h, sup_eq_left.2 h] -lemma interval_swap (a b : α) : [a, b] = [b, a] := -by rw [interval, interval, min_comm, max_comm] +lemma interval_swap (a b : α) : [a, b] = [b, a] := by simp_rw [interval, inf_comm, sup_comm] lemma interval_of_lt (h : a < b) : [a, b] = Icc a b := interval_of_le (le_of_lt h) @@ -55,55 +63,31 @@ interval_of_le (le_of_lt h) lemma interval_of_gt (h : b < a) : [a, b] = Icc b a := interval_of_ge (le_of_lt h) -lemma interval_of_not_le (h : ¬ a ≤ b) : [a, b] = Icc b a := -interval_of_gt (lt_of_not_ge h) +@[simp] lemma interval_self : [a, a] = {a} := by simp [interval] -lemma interval_of_not_ge (h : ¬ b ≤ a) : [a, b] = Icc a b := -interval_of_lt (lt_of_not_ge h) +@[simp] lemma nonempty_interval : [a, b].nonempty := nonempty_Icc.2 inf_le_sup -lemma interval_eq_union : [a, b] = Icc a b ∪ Icc b a := by rw [Icc_union_Icc', max_comm]; refl +lemma Icc_subset_interval : Icc a b ⊆ [a, b] := Icc_subset_Icc inf_le_left le_sup_right +lemma Icc_subset_interval' : Icc b a ⊆ [a, b] := Icc_subset_Icc inf_le_right le_sup_left -lemma mem_interval : a ∈ [b, c] ↔ b ≤ a ∧ a ≤ c ∨ c ≤ a ∧ a ≤ b := by simp [interval_eq_union] +@[simp] lemma left_mem_interval : a ∈ [a, b] := ⟨inf_le_left, le_sup_left⟩ +@[simp] lemma right_mem_interval : b ∈ [a, b] := ⟨inf_le_right, le_sup_right⟩ -@[simp] lemma interval_self : [a, a] = {a} := -set.ext $ by simp [le_antisymm_iff, and_comm] - -@[simp] lemma nonempty_interval : set.nonempty [a, b] := -by { simp only [interval, min_le_iff, le_max_iff, nonempty_Icc], left, left, refl } - -@[simp] lemma left_mem_interval : a ∈ [a, b] := by simp [mem_interval, le_total] -@[simp] lemma right_mem_interval : b ∈ [a, b] := by simp [mem_interval, le_total] - -lemma Icc_subset_interval : Icc a b ⊆ [a, b] := -Icc_subset_Icc (min_le_left _ _) (le_max_right _ _) - -lemma Icc_subset_interval' : Icc b a ⊆ [a, b] := -by { rw interval_swap, apply Icc_subset_interval } - -lemma mem_interval_of_le (ha : a ≤ x) (hb : x ≤ b) : x ∈ [a, b] := -Icc_subset_interval ⟨ha, hb⟩ - -lemma mem_interval_of_ge (hb : b ≤ x) (ha : x ≤ a) : x ∈ [a, b] := -Icc_subset_interval' ⟨hb, ha⟩ - -lemma not_mem_interval_of_lt (ha : c < a) (hb : c < b) : c ∉ [a, b] := -not_mem_Icc_of_lt $ lt_min_iff.mpr ⟨ha, hb⟩ - -lemma not_mem_interval_of_gt (ha : a < c) (hb : b < c) : c ∉ [a, b] := -not_mem_Icc_of_gt $ max_lt_iff.mpr ⟨ha, hb⟩ +lemma mem_interval_of_le (ha : a ≤ x) (hb : x ≤ b) : x ∈ [a, b] := Icc_subset_interval ⟨ha, hb⟩ +lemma mem_interval_of_ge (hb : b ≤ x) (ha : x ≤ a) : x ∈ [a, b] := Icc_subset_interval' ⟨hb, ha⟩ lemma interval_subset_interval (h₁ : a₁ ∈ [a₂, b₂]) (h₂ : b₁ ∈ [a₂, b₂]) : [a₁, b₁] ⊆ [a₂, b₂] := -Icc_subset_Icc (le_min h₁.1 h₂.1) (max_le h₁.2 h₂.2) +Icc_subset_Icc (le_inf h₁.1 h₂.1) (sup_le h₁.2 h₂.2) lemma interval_subset_Icc (ha : a₁ ∈ Icc a₂ b₂) (hb : b₁ ∈ Icc a₂ b₂) : [a₁, b₁] ⊆ Icc a₂ b₂ := -Icc_subset_Icc (le_min ha.1 hb.1) (max_le ha.2 hb.2) +Icc_subset_Icc (le_inf ha.1 hb.1) (sup_le ha.2 hb.2) lemma interval_subset_interval_iff_mem : [a₁, b₁] ⊆ [a₂, b₂] ↔ a₁ ∈ [a₂, b₂] ∧ b₁ ∈ [a₂, b₂] := iff.intro (λh, ⟨h left_mem_interval, h right_mem_interval⟩) (λ h, interval_subset_interval h.1 h.2) -lemma interval_subset_interval_iff_le : - [a₁, b₁] ⊆ [a₂, b₂] ↔ min a₂ b₂ ≤ min a₁ b₁ ∧ max a₁ b₁ ≤ max a₂ b₂ := -by { rw [interval, interval, Icc_subset_Icc_iff], exact min_le_max } +lemma interval_subset_interval_iff_le' : + [a₁, b₁] ⊆ [a₂, b₂] ↔ a₂ ⊓ b₂ ≤ a₁ ⊓ b₁ ∧ a₁ ⊔ b₁ ≤ a₂ ⊔ b₂ := +Icc_subset_Icc_iff inf_le_sup lemma interval_subset_interval_right (h : x ∈ [a, b]) : [x, b] ⊆ [a, b] := interval_subset_interval h right_mem_interval @@ -111,13 +95,20 @@ interval_subset_interval h right_mem_interval lemma interval_subset_interval_left (h : x ∈ [a, b]) : [a, x] ⊆ [a, b] := interval_subset_interval left_mem_interval h -/-- A sort of triangle inequality. -/ -lemma interval_subset_interval_union_interval : [a, c] ⊆ [a, b] ∪ [b, c] := -λ x, by simp only [mem_interval, mem_union]; cases le_total a c; cases le_total x b; tauto +lemma bdd_below_bdd_above_iff_subset_interval (s : set α) : + bdd_below s ∧ bdd_above s ↔ ∃ a b, s ⊆ [a, b] := +bdd_below_bdd_above_iff_subset_Icc.trans + ⟨λ ⟨a, b, h⟩, ⟨a, b, λ x hx, Icc_subset_interval (h hx)⟩, λ ⟨a, b, h⟩, ⟨_, _, h⟩⟩ + +end lattice + +open_locale interval -lemma eq_of_mem_interval_of_mem_interval : a ∈ [b, c] → b ∈ [a, c] → a = b := -by simp_rw mem_interval; rintro (⟨_, _⟩ | ⟨_, _⟩) (⟨_, _⟩ | ⟨_, _⟩); apply le_antisymm; - assumption <|> { exact le_trans ‹_› ‹_› } +section distrib_lattice +variables [distrib_lattice α] {a a₁ a₂ b b₁ b₂ c x : α} + +lemma eq_of_mem_interval_of_mem_interval (ha : a ∈ [b, c]) (hb : b ∈ [a, c]) : a = b := +eq_of_inf_eq_sup_eq (inf_congr_right ha.1 hb.1) $ sup_congr_right ha.2 hb.2 lemma eq_of_mem_interval_of_mem_interval' : b ∈ [a, c] → c ∈ [a, b] → b = c := by simpa only [interval_swap a] using eq_of_mem_interval_of_mem_interval @@ -129,20 +120,40 @@ lemma interval_injective_right (a : α) : injective (λ b, interval b a) := lemma interval_injective_left (a : α) : injective (interval a) := by simpa only [interval_swap] using interval_injective_right a -lemma bdd_below_bdd_above_iff_subset_interval (s : set α) : - bdd_below s ∧ bdd_above s ↔ ∃ a b, s ⊆ [a, b] := -begin - rw [bdd_below_bdd_above_iff_subset_Icc], - split, - { rintro ⟨a, b, h⟩, exact ⟨a, b, λ x hx, Icc_subset_interval (h hx)⟩ }, - { rintro ⟨a, b, h⟩, exact ⟨min a b, max a b, h⟩ } -end +end distrib_lattice + +section linear_order +variables [linear_order α] [linear_order β] {f : α → β} {s : set α} + {a a₁ a₂ b b₁ b₂ c x : α} + +lemma Icc_min_max : Icc (min a b) (max a b) = [a, b] := rfl + +lemma interval_of_not_le (h : ¬ a ≤ b) : [a, b] = Icc b a := interval_of_gt $ lt_of_not_ge h +lemma interval_of_not_ge (h : ¬ b ≤ a) : [a, b] = Icc a b := interval_of_lt $ lt_of_not_ge h + +lemma interval_eq_union : [a, b] = Icc a b ∪ Icc b a := by rw [Icc_union_Icc', max_comm]; refl + +lemma mem_interval : a ∈ [b, c] ↔ b ≤ a ∧ a ≤ c ∨ c ≤ a ∧ a ≤ b := by simp [interval_eq_union] + +lemma not_mem_interval_of_lt (ha : c < a) (hb : c < b) : c ∉ [a, b] := +not_mem_Icc_of_lt $ lt_min_iff.mpr ⟨ha, hb⟩ + +lemma not_mem_interval_of_gt (ha : a < c) (hb : b < c) : c ∉ [a, b] := +not_mem_Icc_of_gt $ max_lt_iff.mpr ⟨ha, hb⟩ + +lemma interval_subset_interval_iff_le : + [a₁, b₁] ⊆ [a₂, b₂] ↔ min a₂ b₂ ≤ min a₁ b₁ ∧ max a₁ b₁ ≤ max a₂ b₂ := +interval_subset_interval_iff_le' + +/-- A sort of triangle inequality. -/ +lemma interval_subset_interval_union_interval : [a, c] ⊆ [a, b] ∪ [b, c] := +λ x, by simp only [mem_interval, mem_union]; cases le_total a c; cases le_total x b; tauto lemma monotone_or_antitone_iff_interval : monotone f ∨ antitone f ↔ ∀ a b c, c ∈ [a, b] → f c ∈ [f a, f b] := begin split, - { rintro (hf | hf) a b c; simp_rw [interval, ←hf.map_min, ←hf.map_max], + { rintro (hf | hf) a b c; simp_rw [←Icc_min_max, ←hf.map_min, ←hf.map_max], exacts [λ hc, ⟨hf hc.1, hf hc.2⟩, λ hc, ⟨hf hc.2, hf hc.1⟩] }, contrapose!, rw not_monotone_not_antitone_iff_exists_le_le, diff --git a/src/data/set/pointwise/interval.lean b/src/data/set/pointwise/interval.lean index d61a6ba39e525..b8f08d0378690 100644 --- a/src/data/set/pointwise/interval.lean +++ b/src/data/set/pointwise/interval.lean @@ -255,19 +255,19 @@ section linear_ordered_add_comm_group variables [linear_ordered_add_comm_group α] (a b c d : α) @[simp] lemma preimage_const_add_interval : (λ x, a + x) ⁻¹' [b, c] = [b - a, c - a] := -by simp only [interval, preimage_const_add_Icc, min_sub_sub_right, max_sub_sub_right] +by simp only [←Icc_min_max, preimage_const_add_Icc, min_sub_sub_right, max_sub_sub_right] @[simp] lemma preimage_add_const_interval : (λ x, x + a) ⁻¹' [b, c] = [b - a, c - a] := by simpa only [add_comm] using preimage_const_add_interval a b c @[simp] lemma preimage_neg_interval : - [a, b] = [-a, -b] := -by simp only [interval, preimage_neg_Icc, min_neg_neg, max_neg_neg] +by simp only [←Icc_min_max, preimage_neg_Icc, min_neg_neg, max_neg_neg] @[simp] lemma preimage_sub_const_interval : (λ x, x - a) ⁻¹' [b, c] = [b + a, c + a] := by simp [sub_eq_add_neg] @[simp] lemma preimage_const_sub_interval : (λ x, a - x) ⁻¹' [b, c] = [a - b, a - c] := -by { rw [interval, interval, preimage_const_sub_Icc], +by { simp_rw [←Icc_min_max, preimage_const_sub_Icc], simp only [sub_eq_add_neg, min_add_add_left, max_add_add_left, min_neg_neg, max_neg_neg], } @[simp] lemma image_const_add_interval : (λ x, a + x) '' [b, c] = [a + b, a + c] := @@ -445,8 +445,8 @@ by simpa only [mul_comm] using preimage_mul_const_Icc_of_neg a b h @[simp] lemma preimage_mul_const_interval (ha : a ≠ 0) (b c : α) : (λ x, x * a) ⁻¹' [b, c] = [b / a, c / a] := (lt_or_gt_of_ne ha).elim - (λ ha, by simp [interval, ha, ha.le, min_div_div_right_of_nonpos, max_div_div_right_of_nonpos]) - (λ (ha : 0 < a), by simp [interval, ha, ha.le, min_div_div_right, max_div_div_right]) + (λ h, by simp [←Icc_min_max, h, h.le, min_div_div_right_of_nonpos, max_div_div_right_of_nonpos]) + (λ (ha : 0 < a), by simp [←Icc_min_max, ha, ha.le, min_div_div_right, max_div_div_right]) @[simp] lemma preimage_const_mul_interval (ha : a ≠ 0) (b c : α) : (λ x, a * x) ⁻¹' [b, c] = [b / a, c / a] := diff --git a/src/measure_theory/integral/interval_integral.lean b/src/measure_theory/integral/interval_integral.lean index 75efc6ed2f3f4..2fc5ac0ba29c6 100644 --- a/src/measure_theory/integral/interval_integral.lean +++ b/src/measure_theory/integral/interval_integral.lean @@ -235,7 +235,7 @@ by rw [integrable_on_Icc_iff_integrable_on_Ioc, integrable_on_Ioc_iff_integrable lemma interval_integrable_iff' [has_no_atoms μ] : interval_integrable f μ a b ↔ integrable_on f (interval a b) μ := -by rw [interval_integrable_iff, interval, interval_oc, integrable_on_Icc_iff_integrable_on_Ioc] +by rw [interval_integrable_iff, ←Icc_min_max, interval_oc, integrable_on_Icc_iff_integrable_on_Ioc] lemma interval_integrable_iff_integrable_Icc_of_le {f : ℝ → E} {a b : ℝ} (hab : a ≤ b) {μ : measure ℝ} [has_no_atoms μ] : diff --git a/src/measure_theory/measure/lebesgue.lean b/src/measure_theory/measure/lebesgue.lean index 96f968b6af59a..2d20d0e8656bf 100644 --- a/src/measure_theory/measure/lebesgue.lean +++ b/src/measure_theory/measure/lebesgue.lean @@ -113,7 +113,7 @@ instance has_no_atoms_volume : has_no_atoms (volume : measure ℝ) := ⟨λ x, volume_singleton⟩ @[simp] lemma volume_interval {a b : ℝ} : volume (interval a b) = of_real (|b - a|) := -by rw [interval, volume_Icc, max_sub_min_eq_abs] +by rw [←Icc_min_max, volume_Icc, max_sub_min_eq_abs] @[simp] lemma volume_Ioi {a : ℝ} : volume (Ioi a) = ∞ := top_unique $ le_of_tendsto' ennreal.tendsto_nat_nhds_top $ λ n, From d97d4fac1fd5788a199f6660a89c21a8e6e6a213 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 14 Dec 2022 05:27:14 +0000 Subject: [PATCH 0037/1291] feat(data/fin/basic): lemmas commuting `cast_add` and `nat_add` (#17917) --- src/data/fin/basic.lean | 16 ++++++++++++++++ 1 file changed, 16 insertions(+) diff --git a/src/data/fin/basic.lean b/src/data/fin/basic.lean index 3c38b7ab48370..bce40a81832ad 100644 --- a/src/data/fin/basic.lean +++ b/src/data/fin/basic.lean @@ -777,6 +777,10 @@ ext rfl cast h (cast_add m' i) = cast_add m i := ext rfl +lemma cast_add_cast_add {m n p : ℕ} (i : fin m) : + cast_add p (cast_add n i) = cast (add_assoc _ _ _).symm (cast_add (n + p) i) := +ext rfl + /-- The cast of the successor is the succesor of the cast. See `fin.succ_cast_eq` for rewriting in the reverse direction. -/ @[simp] lemma cast_succ_eq {n' : ℕ} (i : fin n) (h : n.succ = n'.succ) : @@ -941,6 +945,18 @@ ext rfl cast h (nat_add m' i) = nat_add m i := ext $ (congr_arg (+ (i : ℕ)) (add_right_cancel h) : _) +lemma cast_add_nat_add (p m : ℕ) {n : ℕ} (i : fin n) : + cast_add p (nat_add m i) = cast (add_assoc _ _ _).symm (nat_add m (cast_add p i)) := +ext rfl + +lemma nat_add_cast_add (p m : ℕ) {n : ℕ} (i : fin n) : + nat_add m (cast_add p i) = cast (add_assoc _ _ _) (cast_add p (nat_add m i)) := +ext rfl + +lemma nat_add_nat_add (m n : ℕ) {p : ℕ} (i : fin p) : + nat_add m (nat_add n i) = cast (add_assoc _ _ _) (nat_add (m + n) i) := +ext $ (add_assoc _ _ _).symm + @[simp] lemma cast_nat_add_zero {n n' : ℕ} (i : fin n) (h : 0 + n = n') : cast h (nat_add 0 i) = cast ((zero_add _).symm.trans h) i := ext $ zero_add _ From 198161d833f2c01498c39c266b0b3dbe2c7a8c07 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 14 Dec 2022 05:27:15 +0000 Subject: [PATCH 0038/1291] feat(data/prod/basic): `prod.lex` is trichotomous (#17931) or irreflexive when the base relations are. --- src/data/prod/basic.lean | 24 ++++++++++++++++++++++-- 1 file changed, 22 insertions(+), 2 deletions(-) diff --git a/src/data/prod/basic.lean b/src/data/prod/basic.lean index 226e5aacd1726..023f589c51477 100644 --- a/src/data/prod/basic.lean +++ b/src/data/prod/basic.lean @@ -68,8 +68,9 @@ lemma map_map {ε ζ : Type*} prod.map g g' (prod.map f f' x) = prod.map (g ∘ f) (g' ∘ f') x := rfl -@[simp] theorem mk.inj_iff {a₁ a₂ : α} {b₁ b₂ : β} : (a₁, b₁) = (a₂, b₂) ↔ (a₁ = a₂ ∧ b₁ = b₂) := -⟨prod.mk.inj, by cc⟩ +variables {a a₁ a₂ : α} {b b₁ b₂ : β} + +@[simp] lemma mk.inj_iff : (a₁, b₁) = (a₂, b₂) ↔ a₁ = a₂ ∧ b₁ = b₂ := ⟨prod.mk.inj, by cc⟩ lemma mk.inj_left {α β : Type*} (a : α) : function.injective (prod.mk a : β → α × β) := @@ -79,6 +80,9 @@ lemma mk.inj_right {α β : Type*} (b : β) : function.injective (λ a, prod.mk a b : α → α × β) := by { intros b₁ b₂ h, by simpa only [and_true, eq_self_iff_true, mk.inj_iff] using h } +lemma mk_inj_left : (a, b₁) = (a, b₂) ↔ b₁ = b₂ := (mk.inj_left _).eq_iff +lemma mk_inj_right : (a₁, b) = (a₂, b) ↔ a₁ = a₂ := (mk.inj_right _).eq_iff + lemma ext_iff {p q : α × β} : p = q ↔ p.1 = q.1 ∧ p.2 = q.2 := by rw [← @mk.eta _ _ p, ← @mk.eta _ _ q, mk.inj_iff] @@ -148,6 +152,8 @@ lemma fst_eq_iff : ∀ {p : α × β} {x : α}, p.1 = x ↔ p = (x, p.2) lemma snd_eq_iff : ∀ {p : α × β} {x : β}, p.2 = x ↔ p = (p.1, x) | ⟨a, b⟩ x := by simp +variables {r : α → α → Prop} {s : β → β → Prop} {x y : α × β} + theorem lex_def (r : α → α → Prop) (s : β → β → Prop) {p q : α × β} : prod.lex r s p q ↔ r p.1 q.1 ∨ p.1 = q.1 ∧ s p.2 q.2 := ⟨λ h, by cases h; simp *, @@ -157,6 +163,8 @@ theorem lex_def (r : α → α → Prop) (s : β → β → Prop) by change a = c at e; subst e; exact lex.right _ h end⟩ +lemma lex_iff : lex r s x y ↔ r x.1 y.1 ∨ x.1 = y.1 ∧ s x.2 y.2 := lex_def _ _ + instance lex.decidable [decidable_eq α] (r : α → α → Prop) (s : β → β → Prop) [decidable_rel r] [decidable_rel s] : decidable_rel (prod.lex r s) := @@ -178,6 +186,9 @@ instance is_refl_right {r : α → α → Prop} {s : β → β → Prop} [is_ref is_refl (α × β) (lex r s) := ⟨lex.refl_right _ _⟩ +instance is_irrefl [is_irrefl α r] [is_irrefl β s] : is_irrefl (α × β) (lex r s) := +⟨by rintro ⟨i, a⟩ (⟨_, _, h⟩ | ⟨_, h⟩); exact irrefl _ h⟩ + @[trans] lemma lex.trans {r : α → α → Prop} {s : β → β → Prop} [is_trans α r] [is_trans β s] : ∀ {x y z : α × β}, prod.lex r s x y → prod.lex r s y z → prod.lex r s x z | (x₁, x₂) (y₁, y₂) (z₁, z₂) (lex.left _ _ hxy₁) (lex.left _ _ hyz₁) := @@ -212,6 +223,15 @@ instance is_total_right {r : α → α → Prop} {s : β → β → Prop} [is_tr { exact or.inr (lex.left _ _ hji) } end⟩ +instance is_trichotomous [is_trichotomous α r] [is_trichotomous β s] : + is_trichotomous (α × β) (lex r s) := +⟨λ ⟨i, a⟩ ⟨j, b⟩, begin + obtain hij | rfl | hji := trichotomous_of r i j, + { exact or.inl (lex.left _ _ hij) }, + { exact (trichotomous_of s a b).imp3 (lex.right _) (congr_arg _) (lex.right _) }, + { exact or.inr (or.inr $ lex.left _ _ hji) } +end⟩ + end prod open prod From a2efed67e2728d2886b5d6d40b7950d2a02a8b0b Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Wed, 14 Dec 2022 08:46:47 +0000 Subject: [PATCH 0039/1291] chore(algebra/big_operators/part_enat): drop file (#17934) The only lemma in this file is a version of `nat.cast_sum`. Also fix an instance on `part_enat` so that `nat.cast_sum` works for this type. --- archive/imo/imo2019_q4.lean | 3 +-- src/algebra/big_operators/default.lean | 1 - src/algebra/big_operators/part_enat.lean | 22 ---------------------- src/data/nat/part_enat.lean | 2 +- 4 files changed, 2 insertions(+), 26 deletions(-) delete mode 100644 src/algebra/big_operators/part_enat.lean diff --git a/archive/imo/imo2019_q4.lean b/archive/imo/imo2019_q4.lean index 700aa1be727ae..a4d25853dab7f 100644 --- a/archive/imo/imo2019_q4.lean +++ b/archive/imo/imo2019_q4.lean @@ -5,7 +5,6 @@ Authors: Floris van Doorn -/ import tactic.interval_cases import algebra.big_operators.order -import algebra.big_operators.part_enat import data.nat.multiplicity /-! @@ -36,7 +35,7 @@ begin { suffices : multiplicity 2 (k! : ℤ) = (n * (n - 1) / 2 : ℕ), { rw [← part_enat.coe_lt_coe, ← this], change multiplicity ((2 : ℕ) : ℤ) _ < _, simp_rw [int.coe_nat_multiplicity, multiplicity_two_factorial_lt hk.lt.ne.symm] }, - rw [h, multiplicity.finset.prod prime_2, ← sum_range_id, ← sum_nat_coe_enat], + rw [h, multiplicity.finset.prod prime_2, ← sum_range_id, nat.cast_sum], apply sum_congr rfl, intros i hi, rw [multiplicity_sub_of_gt, multiplicity_pow_self_of_prime prime_2], rwa [multiplicity_pow_self_of_prime prime_2, multiplicity_pow_self_of_prime prime_2, diff --git a/src/algebra/big_operators/default.lean b/src/algebra/big_operators/default.lean index deda0d1890684..7924b69ff6c60 100644 --- a/src/algebra/big_operators/default.lean +++ b/src/algebra/big_operators/default.lean @@ -6,4 +6,3 @@ import algebra.big_operators.ring import algebra.big_operators.pi import algebra.big_operators.finsupp import algebra.big_operators.nat_antidiagonal -import algebra.big_operators.part_enat diff --git a/src/algebra/big_operators/part_enat.lean b/src/algebra/big_operators/part_enat.lean deleted file mode 100644 index e5ac611f40870..0000000000000 --- a/src/algebra/big_operators/part_enat.lean +++ /dev/null @@ -1,22 +0,0 @@ -/- -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 --/ -import algebra.big_operators.basic -import data.nat.part_enat - -/-! -# Big operators in `part_enat` - -A simple lemma about sums in `part_enat`. --/ -open_locale big_operators -variables {α : Type*} - -namespace finset -lemma sum_nat_coe_enat (s : finset α) (f : α → ℕ) : - (∑ x in s, (f x : part_enat)) = (∑ x in s, f x : ℕ) := -(part_enat.coe_hom.map_sum _ _).symm - -end finset diff --git a/src/data/nat/part_enat.lean b/src/data/nat/part_enat.lean index e0407a6320836..b781908972ff0 100644 --- a/src/data/nat/part_enat.lean +++ b/src/data/nat/part_enat.lean @@ -78,7 +78,7 @@ instance : add_comm_monoid part_enat := add_zero := λ x, part.ext' (and_true _) (λ _ _, add_zero _), add_assoc := λ x y z, part.ext' and.assoc (λ _ _, add_assoc _ _ _) } -instance : add_monoid_with_one part_enat := +instance : add_comm_monoid_with_one part_enat := { one := 1, nat_cast := some, nat_cast_zero := rfl, From ae193735714ee12590325bff0a20cba1380fd4d4 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 14 Dec 2022 11:12:14 +0000 Subject: [PATCH 0040/1291] =?UTF-8?q?feat(data/set/pointwise/smul):=20`a?= =?UTF-8?q?=20=E2=80=A2=20(s=20\=20t)=20=3D=20a=20=E2=80=A2=20s=20\=20a=20?= =?UTF-8?q?=E2=80=A2=20t`=20(#17927)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Scalar multiplication distributes over set and symmetric difference. --- src/data/set/pointwise/smul.lean | 53 ++++++++++++++---------- src/group_theory/group_action/group.lean | 20 +++++++-- 2 files changed, 48 insertions(+), 25 deletions(-) diff --git a/src/data/set/pointwise/smul.lean b/src/data/set/pointwise/smul.lean index 6d4f7f8d971b6..77892d5c43e5a 100644 --- a/src/data/set/pointwise/smul.lean +++ b/src/data/set/pointwise/smul.lean @@ -194,24 +194,6 @@ variables {s s₁ s₂ : set α} {t t₁ t₂ : set β} {a : α} {b : β} (⋃ a ∈ t, mul_opposite.op a • s) = s * t := Union_image_right _ -@[to_additive] -lemma smul_set_inter [group α] [mul_action α β] {s t : set β} : - a • (s ∩ t) = a • s ∩ a • t := -(image_inter $ mul_action.injective a).symm - -lemma smul_set_inter₀ [group_with_zero α] [mul_action α β] {s t : set β} (ha : a ≠ 0) : - a • (s ∩ t) = a • s ∩ a • t := -show units.mk0 a ha • _ = _, from smul_set_inter - -@[simp, to_additive] -lemma smul_set_univ [group α] [mul_action α β] {a : α} : a • (univ : set β) = univ := -eq_univ_of_forall $ λ b, ⟨a⁻¹ • b, trivial, smul_inv_smul _ _⟩ - -@[simp, to_additive] -lemma smul_univ [group α] [mul_action α β] {s : set α} (hs : s.nonempty) : - s • (univ : set β) = univ := -let ⟨a, ha⟩ := hs in eq_univ_of_forall $ λ b, ⟨a, a⁻¹ • b, ha, trivial, smul_inv_smul _ _⟩ - @[to_additive] theorem range_smul_range {ι κ : Type*} [has_smul α β] (b : ι → α) (c : κ → β) : range b • range c = range (λ p : ι × κ, b p.1 • c p.2) := @@ -490,6 +472,21 @@ lemma subset_set_smul_iff : A ⊆ a • B ↔ a⁻¹ • A ⊆ B := iff.symm $ (image_subset_iff).trans $ iff.symm $ iff_of_eq $ congr_arg _ $ image_equiv_eq_preimage_symm _ $ mul_action.to_perm _ +@[to_additive] lemma smul_set_inter : a • (s ∩ t) = a • s ∩ a • t := +(image_inter $ mul_action.injective a).symm + +@[to_additive] lemma smul_set_sdiff : a • (s \ t) = a • s \ a • t := +image_diff (mul_action.injective a) _ _ + +@[to_additive] lemma smul_set_symm_diff : a • (s ∆ t) = (a • s) ∆ (a • t) := +image_symm_diff (mul_action.injective a) _ _ + +@[simp, to_additive] lemma smul_set_univ : a • (univ : set β) = univ := +image_univ_of_surjective $ mul_action.surjective a + +@[simp, to_additive] lemma smul_univ {s : set α} (hs : s.nonempty) : s • (univ : set β) = univ := +let ⟨a, ha⟩ := hs in eq_univ_of_forall $ λ b, ⟨a, a⁻¹ • b, ha, trivial, smul_inv_smul _ _⟩ + @[to_additive] lemma smul_inter_ne_empty_iff {s t : set α} {x : α} : x • s ∩ t ≠ ∅ ↔ ∃ a b, (a ∈ t ∧ b ∈ s) ∧ a * b⁻¹ = x := @@ -534,7 +531,7 @@ by simp_rw [← Union_set_of, ← Union_inv_smul, ← preimage_smul, preimage] end group section group_with_zero -variables [group_with_zero α] [mul_action α β] {s : set α} {a : α} +variables [group_with_zero α] [mul_action α β] {s t : set β} {a : α} @[simp] lemma smul_mem_smul_set_iff₀ (ha : a ≠ 0) (A : set β) (x : β) : a • x ∈ a • A ↔ x ∈ A := @@ -564,12 +561,24 @@ show units.mk0 a ha • _ ⊆ _ ↔ _, from set_smul_subset_iff lemma subset_set_smul_iff₀ (ha : a ≠ 0) {A B : set β} : A ⊆ a • B ↔ a⁻¹ • A ⊆ B := show _ ⊆ units.mk0 a ha • _ ↔ _, from subset_set_smul_iff -lemma smul_univ₀ (hs : ¬ s ⊆ 0) : s • (univ : set β) = univ := +lemma smul_set_inter₀ (ha : a ≠ 0) : a • (s ∩ t) = a • s ∩ a • t := +show units.mk0 a ha • _ = _, from smul_set_inter + +lemma smul_set_sdiff₀ (ha : a ≠ 0) : a • (s \ t) = a • s \ a • t := +image_diff (mul_action.injective₀ ha) _ _ + +lemma smul_set_symm_diff₀ (ha : a ≠ 0) : a • (s ∆ t) = (a • s) ∆ (a • t) := +image_symm_diff (mul_action.injective₀ ha) _ _ + +lemma smul_set_univ₀ (ha : a ≠ 0) : a • (univ : set β) = univ := +image_univ_of_surjective $ mul_action.surjective₀ ha + +lemma smul_univ₀ {s : set α} (hs : ¬ s ⊆ 0) : s • (univ : set β) = univ := let ⟨a, ha, ha₀⟩ := not_subset.1 hs in eq_univ_of_forall $ λ b, ⟨a, a⁻¹ • b, ha, trivial, smul_inv_smul₀ ha₀ _⟩ -lemma smul_set_univ₀ (ha : a ≠ 0) : a • (univ : set β) = univ := -eq_univ_of_forall $ λ b, ⟨a⁻¹ • b, trivial, smul_inv_smul₀ ha _⟩ +lemma smul_univ₀' {s : set α} (hs : s.nontrivial) : s • (univ : set β) = univ := +smul_univ₀ hs.not_subset_singleton end group_with_zero diff --git a/src/group_theory/group_action/group.lean b/src/group_theory/group_action/group.lean index 916cc6f5542f8..8747e9d145e65 100644 --- a/src/group_theory/group_action/group.lean +++ b/src/group_theory/group_action/group.lean @@ -12,6 +12,8 @@ import group_theory.group_action.units This file contains lemmas about `smul` on `group_with_zero`, and `group`. -/ +open function + universes u v w variables {α : Type u} {β : Type v} {γ : Type w} @@ -105,12 +107,15 @@ by { cases p; simp [smul_pow, smul_inv] } commute (r • a) b ↔ commute a b := by rw [commute.symm_iff, commute.smul_right_iff, commute.symm_iff] -@[to_additive] protected lemma mul_action.bijective (g : α) : function.bijective (λ b : β, g • b) := +@[to_additive] protected lemma mul_action.bijective (g : α) : bijective ((•) g : β → β) := (mul_action.to_perm g).bijective -@[to_additive] protected lemma mul_action.injective (g : α) : function.injective (λ b : β, g • b) := +@[to_additive] protected lemma mul_action.injective (g : α) : injective ((•) g : β → β) := (mul_action.bijective g).injective +@[to_additive] protected lemma mul_action.surjective (g : α) : surjective ((•) g : β → β) := +(mul_action.bijective g).surjective + @[to_additive] lemma smul_left_cancel (g : α) {x y : β} (h : g • x = g • y) : x = y := mul_action.injective g h @@ -129,7 +134,7 @@ instance cancel_monoid_with_zero.to_has_faithful_smul [cancel_monoid_with_zero ⟨λ x y h, mul_left_injective₀ one_ne_zero (h 1)⟩ section gwz -variables [group_with_zero α] [mul_action α β] +variables [group_with_zero α] [mul_action α β] {a : α} @[simp] lemma inv_smul_smul₀ {c : α} (hc : c ≠ 0) (x : β) : c⁻¹ • c • x = x := @@ -155,6 +160,15 @@ commute.smul_right_iff (units.mk0 c hc) commute (c • a) b ↔ commute a b := commute.smul_left_iff (units.mk0 c hc) +protected lemma mul_action.bijective₀ (ha : a ≠ 0) : bijective ((•) a : β → β) := +mul_action.bijective $ units.mk0 a ha + +protected lemma mul_action.injective₀ (ha : a ≠ 0) : injective ((•) a : β → β) := +(mul_action.bijective₀ ha).injective + +protected lemma mul_action.surjective₀ (ha : a ≠ 0) : surjective ((•) a : β → β) := +(mul_action.bijective₀ ha).surjective + end gwz end mul_action From ae0d10f6c1e0d389391d867d5ca486c17a2873eb Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Mar=C3=ADa=20In=C3=A9s=20de=20Frutos-Fern=C3=A1ndez?= Date: Wed, 14 Dec 2022 14:02:52 +0000 Subject: [PATCH 0041/1291] feat(ring_theory/dedekind_domain/factorization): add factorization lemmas (#17915) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This is the first in a sequence of PRs to prove that every nonzero fractional ideal of a Dedekind domain `R` can be factored as a finprod `∏_v v^{n_v}` over the maximal ideals of `R`, where the exponents `n_v` are integers, and to provide related API. In this PR we prove the analogous statement for ideals of `R`. --- .../dedekind_domain/factorization.lean | 185 ++++++++++++++++++ 1 file changed, 185 insertions(+) create mode 100644 src/ring_theory/dedekind_domain/factorization.lean diff --git a/src/ring_theory/dedekind_domain/factorization.lean b/src/ring_theory/dedekind_domain/factorization.lean new file mode 100644 index 0000000000000..e2aa140725183 --- /dev/null +++ b/src/ring_theory/dedekind_domain/factorization.lean @@ -0,0 +1,185 @@ +/- +Copyright (c) 2022 María Inés de Frutos-Fernández. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: María Inés de Frutos-Fernández +-/ +import ring_theory.dedekind_domain.ideal +/-! +# Factorization of ideals of Dedekind domains +Every nonzero ideal `I` of a Dedekind domain `R` can be factored as a product `∏_v v^{n_v}` over the +maximal ideals of `R`, where the exponents `n_v` are natural numbers. +TODO: Extend the results in this file to fractional ideals of `R`. +## Main results +- `ideal.finite_factors` : Only finitely many maximal ideals of `R` divide a given nonzero ideal. +- `ideal.finprod_height_one_spectrum_factorization` : The ideal `I` equals the finprod + `∏_v v^(val_v(I))`,where `val_v(I)` denotes the multiplicity of `v` in the factorization of `I` + and `v` runs over the maximal ideals of `R`. +## Tags +dedekind domain, ideal, factorization +-/ + +noncomputable theory +open_locale big_operators classical non_zero_divisors + +open set function unique_factorization_monoid is_dedekind_domain + is_dedekind_domain.height_one_spectrum + +/-! ### Factorization of ideals of Dedekind domains -/ + +variables {R : Type*} [comm_ring R] [is_domain R] [is_dedekind_domain R] {K : Type*} [field K] + [algebra R K] [is_fraction_ring R K] (v : height_one_spectrum R) + +/-- Given a maximal ideal `v` and an ideal `I` of `R`, `max_pow_dividing` returns the maximal + power of `v` dividing `I`. -/ +def is_dedekind_domain.height_one_spectrum.max_pow_dividing (I : ideal R) : ideal R := +v.as_ideal^(associates.mk v.as_ideal).count (associates.mk I).factors + +/-- Only finitely many maximal ideals of `R` divide a given nonzero ideal. -/ +lemma ideal.finite_factors {I : ideal R} (hI : I ≠ 0) : + {v : height_one_spectrum R | v.as_ideal ∣ I}.finite := +begin + rw [← set.finite_coe_iff, set.coe_set_of], + haveI h_fin := fintype_subtype_dvd I hI, + refine finite.of_injective (λ v, (⟨(v : height_one_spectrum R).as_ideal, v.2⟩ : {x // x ∣ I})) _, + intros v w hvw, + simp only at hvw, + exact subtype.coe_injective ((height_one_spectrum.ext_iff ↑v ↑w).mpr hvw) +end + +/-- For every nonzero ideal `I` of `v`, there are finitely many maximal ideals `v` such that the + multiplicity of `v` in the factorization of `I`, denoted `val_v(I)`, is nonzero. -/ +lemma associates.finite_factors {I : ideal R} (hI : I ≠ 0) : + ∀ᶠ (v : height_one_spectrum R) in filter.cofinite, + ((associates.mk v.as_ideal).count (associates.mk I).factors : ℤ) = 0 := +begin + have h_supp : {v : height_one_spectrum R | + ¬((associates.mk v.as_ideal).count (associates.mk I).factors : ℤ) = 0} = + {v : height_one_spectrum R | v.as_ideal ∣ I}, + { ext v, + simp_rw int.coe_nat_eq_zero, + exact associates.count_ne_zero_iff_dvd hI v.irreducible, }, + rw [filter.eventually_cofinite, h_supp], + exact ideal.finite_factors hI, +end + +namespace ideal + +/-- For every nonzero ideal `I` of `v`, there are finitely many maximal ideals `v` such that + `v^(val_v(I))` is not the unit ideal. -/ +lemma finite_mul_support {I : ideal R} (hI : I ≠ 0) : + (mul_support (λ (v : height_one_spectrum R), v.max_pow_dividing I)).finite := +begin + have h_subset : {v : height_one_spectrum R | v.max_pow_dividing I ≠ 1} ⊆ + {v : height_one_spectrum R | + ((associates.mk v.as_ideal).count (associates.mk I).factors : ℤ) ≠ 0}, + { intros v hv h_zero, + have hv' : v.max_pow_dividing I = 1, + { rw [is_dedekind_domain.height_one_spectrum.max_pow_dividing, int.coe_nat_eq_zero.mp h_zero, + pow_zero _] }, + exact hv hv', }, + exact finite.subset (filter.eventually_cofinite.mp (associates.finite_factors hI)) h_subset, +end + +/-- For every nonzero ideal `I` of `v`, there are finitely many maximal ideals `v` such that +`v^(val_v(I))`, regarded as a fractional ideal, is not `(1)`. -/ +lemma finite_mul_support_coe {I : ideal R} (hI : I ≠ 0) : + (mul_support (λ (v : height_one_spectrum R), + (v.as_ideal : fractional_ideal R⁰ K) ^ + ((associates.mk v.as_ideal).count (associates.mk I).factors : ℤ))).finite := +begin + rw mul_support, + simp_rw [ne.def, zpow_coe_nat, ← fractional_ideal.coe_ideal_pow, + fractional_ideal.coe_ideal_eq_one_iff], + exact finite_mul_support hI, +end + +/-- For every nonzero ideal `I` of `v`, there are finitely many maximal ideals `v` such that +`v^-(val_v(I))` is not the unit ideal. -/ +lemma finite_mul_support_inv {I : ideal R} (hI : I ≠ 0) : + (mul_support (λ (v : height_one_spectrum R), + (v.as_ideal : fractional_ideal R⁰ K) ^ + -((associates.mk v.as_ideal).count (associates.mk I).factors : ℤ))).finite := +begin + rw mul_support, + simp_rw [zpow_neg, ne.def, inv_eq_one], + exact finite_mul_support_coe hI, +end + +/-- For every nonzero ideal `I` of `v`, `v^(val_v(I) + 1)` does not divide `∏_v v^(val_v(I))`. -/ +lemma finprod_not_dvd (I : ideal R) (hI : I ≠ 0) : + ¬ (v.as_ideal) ^ ((associates.mk v.as_ideal).count (associates.mk I).factors + 1) ∣ + (∏ᶠ (v : height_one_spectrum R), v.max_pow_dividing I) := +begin + have hf := finite_mul_support hI, + have h_ne_zero : v.max_pow_dividing I ≠ 0 := pow_ne_zero _ v.ne_bot, + rw [← mul_finprod_cond_ne v hf, pow_add, pow_one, finprod_cond_ne _ _ hf], + intro h_contr, + have hv_prime : prime v.as_ideal := ideal.prime_of_is_prime v.ne_bot v.is_prime, + obtain ⟨w, hw, hvw'⟩ := + prime.exists_mem_finset_dvd hv_prime ((mul_dvd_mul_iff_left h_ne_zero).mp h_contr), + have hw_prime : prime w.as_ideal := ideal.prime_of_is_prime w.ne_bot w.is_prime, + have hvw := prime.dvd_of_dvd_pow hv_prime hvw', + rw [prime.dvd_prime_iff_associated hv_prime hw_prime, associated_iff_eq] at hvw, + exact (finset.mem_erase.mp hw).1 (height_one_spectrum.ext w v (eq.symm hvw)), +end + +end ideal + +lemma associates.finprod_ne_zero (I : ideal R) : + associates.mk (∏ᶠ (v : height_one_spectrum R), v.max_pow_dividing I) ≠ 0 := +begin + rw [associates.mk_ne_zero, finprod_def], + split_ifs, + { rw finset.prod_ne_zero_iff, + intros v hv, + apply pow_ne_zero _ v.ne_bot, }, + { exact one_ne_zero, } +end + +namespace ideal + +/-- The multiplicity of `v` in `∏_v v^(val_v(I))` equals `val_v(I)`. -/ +lemma finprod_count (I : ideal R) (hI : I ≠ 0) : (associates.mk v.as_ideal).count + (associates.mk (∏ᶠ (v : height_one_spectrum R), v.max_pow_dividing I)).factors = + (associates.mk v.as_ideal).count (associates.mk I).factors := +begin + have h_ne_zero := associates.finprod_ne_zero I, + have hv : irreducible (associates.mk v.as_ideal) := v.associates_irreducible, + have h_dvd := finprod_mem_dvd v (ideal.finite_mul_support hI), + have h_not_dvd := ideal.finprod_not_dvd v I hI, + simp only [is_dedekind_domain.height_one_spectrum.max_pow_dividing] at h_dvd h_ne_zero h_not_dvd, + rw [← associates.mk_dvd_mk, associates.dvd_eq_le, associates.mk_pow, + associates.prime_pow_dvd_iff_le h_ne_zero hv] at h_dvd h_not_dvd, + rw not_le at h_not_dvd, + apply nat.eq_of_le_of_lt_succ h_dvd h_not_dvd, +end + +/-- The ideal `I` equals the finprod `∏_v v^(val_v(I))`. -/ +lemma finprod_height_one_spectrum_factorization (I : ideal R) (hI : I ≠ 0) : + ∏ᶠ (v : height_one_spectrum R), v.max_pow_dividing I = I := +begin + rw [← associated_iff_eq, ← associates.mk_eq_mk_iff_associated], + apply associates.eq_of_eq_counts, + { apply associates.finprod_ne_zero I }, + { apply associates.mk_ne_zero.mpr hI }, + intros v hv, + obtain ⟨J, hJv⟩ := associates.exists_rep v, + rw [← hJv, associates.irreducible_mk] at hv, + rw ← hJv, + apply ideal.finprod_count ⟨J, ideal.is_prime_of_prime (irreducible_iff_prime.mp hv), + irreducible.ne_zero hv⟩ I hI, +end + +/-- The ideal `I` equals the finprod `∏_v v^(val_v(I))`, when both sides are regarded as fractional +ideals of `R`. -/ +lemma finprod_height_one_spectrum_factorization_coe (I : ideal R) (hI : I ≠ 0) : + ∏ᶠ (v : height_one_spectrum R), (v.as_ideal : fractional_ideal R⁰ K) ^ + ((associates.mk v.as_ideal).count (associates.mk I).factors : ℤ) = I := +begin + conv_rhs { rw ← ideal.finprod_height_one_spectrum_factorization I hI }, + rw fractional_ideal.coe_ideal_finprod R⁰ K (le_refl _), + simp_rw [is_dedekind_domain.height_one_spectrum.max_pow_dividing, fractional_ideal.coe_ideal_pow, + zpow_coe_nat], +end + +end ideal From 90367774bb3afc7bdc1e9acbc770970042378306 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 14 Dec 2022 14:02:53 +0000 Subject: [PATCH 0042/1291] fix(tactic/assert_exists): use more unique names (#17916) Previously we were relying on `tactic.mk_fresh_name` being globally unique, but this turned out to not be the case. Now we also combine the declaration name / instance string with the fresh name, to reduce the chance of conflicts. --- src/tactic/assert_exists.lean | 5 +++-- 1 file changed, 3 insertions(+), 2 deletions(-) diff --git a/src/tactic/assert_exists.lean b/src/tactic/assert_exists.lean index 876df3c4da70a..3d120fe89243a 100644 --- a/src/tactic/assert_exists.lean +++ b/src/tactic/assert_exists.lean @@ -62,7 +62,7 @@ do ff ← succeeds (get_decl decl) | fail format!"Declaration {decl} is not allowed to exist in this file.", n ← tactic.mk_fresh_name, - let marker := (`assert_not_exists._checked).append n, + let marker := (`assert_not_exists._checked).append (decl.append n), add_decl (declaration.defn marker [] `(name) `(decl) default tt), pure () @@ -126,7 +126,8 @@ do match i with | none := do n ← tactic.mk_fresh_name, - let marker := (`assert_no_instance._checked).append n, + e_str ← to_string <$> pp e, + let marker := ((`assert_no_instance._checked).mk_string e_str).append n, et ← infer_type e, tt ← succeeds (get_decl marker) | add_decl From 1897d7dedd17ea0c3dd16117043acdf9afbfd56a Mon Sep 17 00:00:00 2001 From: Oliver Nash Date: Wed, 14 Dec 2022 17:11:36 +0000 Subject: [PATCH 0043/1291] feat(measure_theory/covering/liminf_limsup): add lemma `blimsup_cthickening_mul_ae_eq` (#17803) --- docs/references.bib | 16 ++ .../covering/liminf_limsup.lean | 219 ++++++++++++++++++ src/measure_theory/measure/doubling.lean | 23 +- .../measure/measure_space_def.lean | 4 + src/order/filter/basic.lean | 9 + src/order/liminf_limsup.lean | 46 ++++ .../metric_space/hausdorff_distance.lean | 33 ++- src/topology/order/basic.lean | 35 +++ 8 files changed, 383 insertions(+), 2 deletions(-) create mode 100644 src/measure_theory/covering/liminf_limsup.lean diff --git a/docs/references.bib b/docs/references.bib index 2baa09a41bcdb..84b6f8ab0b2a8 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -483,6 +483,22 @@ @Book{ cassels1967algebraic mrclass = {00.04 (10.00)} } +@Article{ cassels1950, + author = {Cassels, J. W. S.}, + title = {Some metrical theorems in {D}iophantine approximation. {I}}, + journal = {Proc. Cambridge Philos. Soc.}, + fjournal = {Proceedings of the Cambridge Philosophical Society}, + volume = {46}, + year = {1950}, + pages = {209--218}, + issn = {0008-1981}, + mrclass = {10.0X}, + mrnumber = {36787}, + mrreviewer = {P. Erd\H{o}s}, + doi = {10.1017/s0305004100025676}, + url = {https://doi.org/10.1017/s0305004100025676}, +} + @InProceedings{ Chou1994, author = {Chou, Ching-Tsun}, booktitle = {Higher Order Logic Theorem Proving and Its Applications}, diff --git a/src/measure_theory/covering/liminf_limsup.lean b/src/measure_theory/covering/liminf_limsup.lean new file mode 100644 index 0000000000000..498f872440ae8 --- /dev/null +++ b/src/measure_theory/covering/liminf_limsup.lean @@ -0,0 +1,219 @@ +/- +Copyright (c) 2022 Oliver Nash. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Oliver Nash +-/ +import measure_theory.covering.density_theorem + +/-! +# Liminf, limsup, and doubling measures. + +This file is a place to collect lemmas about liminf and limsup for subsets of a metric space +carrying a doubling measure. + +## Main results: + + * `blimsup_cthickening_mul_ae_eq`: the limsup of the closed thickening of a sequence of subsets + of a metric space is unchanged almost everywhere for a doubling measure if the sequence of + distances is multiplied by a positive scale factor. This is a generalisation of a result of + Cassels, appearing as Lemma 9 on page 217 of + [J.W.S. Cassels, *Some metrical theorems in Diophantine approximation. I*](cassels1950). + +-/ + +open set filter metric measure_theory +open_locale nnreal ennreal topological_space + +variables {α : Type*} [metric_space α] [sigma_compact_space α] [measurable_space α] [borel_space α] +variables (μ : measure α) [is_locally_finite_measure μ] [is_doubling_measure μ] + +/-- This is really an auxiliary result en route to `blimsup_cthickening_ae_le_of_eventually_mul_le` +(which is itself an auxiliary result en route to `blimsup_cthickening_mul_ae_eq`). + +NB: The `set : α` type ascription is present because of issue #16932 on GitHub. -/ +lemma blimsup_cthickening_ae_le_of_eventually_mul_le_aux + (p : ℕ → Prop) {s : ℕ → set α} (hs : ∀ i, is_closed (s i)) + {r₁ r₂ : ℕ → ℝ} (hr : tendsto r₁ at_top (𝓝[>] 0)) (hrp : 0 ≤ r₁) + {M : ℝ} (hM : 0 < M) (hM' : M < 1) (hMr : ∀ᶠ i in at_top, M * r₁ i ≤ r₂ i) : + (blimsup (λ i, cthickening (r₁ i) (s i)) at_top p : set α) ≤ᵐ[μ] + (blimsup (λ i, cthickening (r₂ i) (s i)) at_top p : set α) := +begin + /- Sketch of proof: + + Assume that `p` is identically true for simplicity. Let `Y₁ i = cthickening (r₁ i) (s i)`, define + `Y₂` similarly except using `r₂`, and let `(Z i) = ⋃_{j ≥ i} (Y₂ j)`. Our goal is equivalent to + showing that `μ ((limsup Y₁) \ (Z i)) = 0` for all `i`. + + Assume for contradiction that `μ ((limsup Y₁) \ (Z i)) ≠ 0` for some `i` and let + `W = (limsup Y₁) \ (Z i)`. Apply Lebesgue's density theorem to obtain a point `d` in `W` of + density `1`. Since `d ∈ limsup Y₁`, there is a subsequence of `j ↦ Y₁ j`, indexed by + `f 0 < f 1 < ...`, such that `d ∈ Y₁ (f j)` for all `j`. For each `j`, we may thus choose + `w j ∈ s (f j)` such that `d ∈ B j`, where `B j = closed_ball (w j) (r₁ (f j))`. Note that + since `d` has density one, `μ (W ∩ (B j)) / μ (B j) → 1`. + + We obtain our contradiction by showing that there exists `η < 1` such that + `μ (W ∩ (B j)) / μ (B j) ≤ η` for sufficiently large `j`. In fact we claim that `η = 1 - C⁻¹` + is such a value where `C` is the scaling constant of `M⁻¹` for the doubling measure `μ`. + + To prove the claim, let `b j = closed_ball (w j) (M * r₁ (f j))` and for given `j` consider the + sets `b j` and `W ∩ (B j)`. These are both subsets of `B j` and are disjoint for large enough `j` + since `M * r₁ j ≤ r₂ j` and thus `b j ⊆ Z i ⊆ Wᶜ`. We thus have: + `μ (b j) + μ (W ∩ (B j)) ≤ μ (B j)`. Combining this with `μ (B j) ≤ C * μ (b j)` we obtain + the required inequality. -/ + set Y₁ : ℕ → set α := λ i, cthickening (r₁ i) (s i), + set Y₂ : ℕ → set α := λ i, cthickening (r₂ i) (s i), + let Z : ℕ → set α := λ i, ⋃ j (h : p j ∧ i ≤ j), Y₂ j, + suffices : ∀ i, μ (at_top.blimsup Y₁ p \ Z i) = 0, + { rwa [ae_le_set, @blimsup_eq_infi_bsupr_of_nat _ _ _ Y₂, infi_eq_Inter, diff_Inter, + measure_Union_null_iff], }, + intros, + set W := at_top.blimsup Y₁ p \ Z i, + by_contra contra, + obtain ⟨d, hd, hd'⟩ : ∃ d, d ∈ W ∧ ∀ {ι : Type*} {l : filter ι} (w : ι → α) (δ : ι → ℝ), + tendsto δ l (𝓝[>] 0) → (∀ᶠ j in l, d ∈ closed_ball (w j) (2 * δ j)) → + tendsto (λ j, μ (W ∩ closed_ball (w j) (δ j)) / μ (closed_ball (w j) (δ j))) l (𝓝 1) := + measure.exists_mem_of_measure_ne_zero_of_ae contra + (is_doubling_measure.ae_tendsto_measure_inter_div μ W 2), + replace hd : d ∈ blimsup Y₁ at_top p := ((mem_diff _).mp hd).1, + obtain ⟨f : ℕ → ℕ, hf⟩ := exists_forall_mem_of_has_basis_mem_blimsup' at_top_basis hd, + simp only [forall_and_distrib] at hf, + obtain ⟨hf₀ : ∀ j, d ∈ cthickening (r₁ (f j)) (s (f j)), hf₁, hf₂ : ∀ j, j ≤ f j⟩ := hf, + have hf₃ : tendsto f at_top at_top := + tendsto_at_top_at_top.mpr (λ j, ⟨f j, λ i hi, (hf₂ j).trans (hi.trans $ hf₂ i)⟩), + replace hr : tendsto (r₁ ∘ f) at_top (𝓝[>] 0) := hr.comp hf₃, + replace hMr : ∀ᶠ j in at_top, M * r₁ (f j) ≤ r₂ (f j) := hf₃.eventually hMr, + replace hf₀ : ∀ j, ∃ (w ∈ s (f j)), d ∈ closed_ball w (2 * r₁ (f j)), + { intros j, + specialize hrp (f j), + rw pi.zero_apply at hrp, + rcases eq_or_lt_of_le hrp with hr0 | hrp', + { specialize hf₀ j, + rw [← hr0, cthickening_zero, (hs (f j)).closure_eq] at hf₀, + exact ⟨d, hf₀, by simp [← hr0]⟩, }, + { exact mem_Union₂.mp (cthickening_subset_Union_closed_ball_of_lt (s (f j)) (by positivity) + (lt_two_mul_self hrp') (hf₀ j)), }, }, + choose w hw hw' using hf₀, + let C := is_doubling_measure.scaling_constant_of μ M⁻¹, + have hC : 0 < C := + lt_of_lt_of_le zero_lt_one (is_doubling_measure.one_le_scaling_constant_of μ M⁻¹), + suffices : ∃ (η < (1 : ℝ≥0)), ∀ᶠ j in at_top, + μ (W ∩ closed_ball (w j) (r₁ (f j))) / μ (closed_ball (w j) (r₁ (f j))) ≤ η, + { obtain ⟨η, hη, hη'⟩ := this, + replace hη' : 1 ≤ η := by simpa only [ennreal.one_le_coe_iff] using + le_of_tendsto (hd' w (λ j, r₁ (f j)) hr $ eventually_of_forall hw') hη', + exact (lt_self_iff_false _).mp (lt_of_lt_of_le hη hη'), }, + refine ⟨1 - C⁻¹, tsub_lt_self zero_lt_one (nnreal.inv_pos.mpr hC), _⟩, + replace hC : C ≠ 0 := ne_of_gt hC, + let b : ℕ → set α := λ j, closed_ball (w j) (M * r₁ (f j)), + let B : ℕ → set α := λ j, closed_ball (w j) (r₁ (f j)), + have h₁ : ∀ j, b j ⊆ B j := + λ j, closed_ball_subset_closed_ball (mul_le_of_le_one_left (hrp (f j)) hM'.le), + have h₂ : ∀ j, W ∩ B j ⊆ B j := λ j, inter_subset_right W (B j), + have h₃ : ∀ᶠ j in at_top, disjoint (b j) (W ∩ B j), + { apply hMr.mp, + rw eventually_at_top, + refine ⟨i, λ j hj hj', disjoint.inf_right (B j) $ disjoint.inf_right' (blimsup Y₁ at_top p) _⟩, + change disjoint (b j) (Z i)ᶜ, + rw disjoint_compl_right_iff_subset, + refine (closed_ball_subset_cthickening (hw j) (M * r₁ (f j))).trans + ((cthickening_mono hj' _).trans (λ a ha, _)), + simp only [mem_Union, exists_prop], + exact ⟨f j, ⟨hf₁ j, hj.le.trans (hf₂ j)⟩, ha⟩, }, + have h₄ : ∀ᶠ j in at_top, μ (B j) ≤ C * μ (b j) := + (hr.eventually (is_doubling_measure.eventually_measure_le_scaling_constant_mul' + μ M hM)).mono (λ j hj, hj (w j)), + refine (h₃.and h₄).mono (λ j hj₀, _), + change μ (W ∩ B j) / μ (B j) ≤ ↑(1 - C⁻¹), + rcases eq_or_ne (μ (B j)) ∞ with hB | hB, { simp [hB], }, + apply ennreal.div_le_of_le_mul, + rw [with_top.coe_sub, ennreal.coe_one, ennreal.sub_mul (λ _ _, hB), one_mul], + replace hB : ↑C⁻¹ * μ (B j) ≠ ∞, + { refine ennreal.mul_ne_top _ hB, + rwa [ennreal.coe_inv hC, ne.def, ennreal.inv_eq_top, ennreal.coe_eq_zero], }, + obtain ⟨hj₁ : disjoint (b j) (W ∩ B j), hj₂ : μ (B j) ≤ C * μ (b j)⟩ := hj₀, + replace hj₂ : ↑C⁻¹ * μ (B j) ≤ μ (b j), + { rw [ennreal.coe_inv hC, ← ennreal.div_eq_inv_mul], + exact ennreal.div_le_of_le_mul' hj₂, }, + have hj₃ : ↑C⁻¹ * μ (B j) + μ (W ∩ B j) ≤ μ (B j), + { refine le_trans (add_le_add_right hj₂ _) _, + rw ← measure_union' hj₁ measurable_set_closed_ball, + exact measure_mono (union_subset (h₁ j) (h₂ j)), }, + replace hj₃ := tsub_le_tsub_right hj₃ (↑C⁻¹ * μ (B j)), + rwa ennreal.add_sub_cancel_left hB at hj₃, +end + +/-- This is really an auxiliary result en route to `blimsup_cthickening_mul_ae_eq`. + +NB: The `set : α` type ascription is present because of issue #16932 on GitHub. -/ +lemma blimsup_cthickening_ae_le_of_eventually_mul_le + (p : ℕ → Prop) {s : ℕ → set α} {M : ℝ} (hM : 0 < M) + {r₁ r₂ : ℕ → ℝ} (hr : tendsto r₁ at_top (𝓝[>] 0)) (hMr : ∀ᶠ i in at_top, M * r₁ i ≤ r₂ i) : + (blimsup (λ i, cthickening (r₁ i) (s i)) at_top p : set α) ≤ᵐ[μ] + (blimsup (λ i, cthickening (r₂ i) (s i)) at_top p : set α) := +begin + let R₁ := λ i, max 0 (r₁ i), + let R₂ := λ i, max 0 (r₂ i), + have hRp : 0 ≤ R₁ := λ i, le_max_left 0 (r₁ i), + replace hMr : ∀ᶠ i in at_top, M * R₁ i ≤ R₂ i, + { refine hMr.mono (λ i hi, _), + rw [mul_max_of_nonneg _ _ hM.le, mul_zero], + exact max_le_max (le_refl 0) hi, }, + simp_rw [← cthickening_max_zero (r₁ _), ← cthickening_max_zero (r₂ _)], + cases le_or_lt 1 M with hM' hM', + { apply has_subset.subset.eventually_le, + change _ ≤ _, + refine mono_blimsup' (hMr.mono $ λ i hi, cthickening_mono _ (s i)), + exact (le_mul_of_one_le_left (hRp i) hM').trans hi, }, + { simp only [← @cthickening_closure _ _ _ (s _)], + have hs : ∀ i, is_closed (closure (s i)) := λ i, is_closed_closure, + exact blimsup_cthickening_ae_le_of_eventually_mul_le_aux + μ p hs (tendsto_nhds_max_right hr) hRp hM hM' hMr, }, +end + +/-- Given a sequence of subsets `sᵢ` of a metric space, together with a sequence of radii `rᵢ` +such that `rᵢ → 0`, the set of points which belong to infinitely many of the closed +`rᵢ`-thickenings of `sᵢ` is unchanged almost everywhere for a doubling measure if the `rᵢ` are all +scaled by a positive constant. + +This lemma is a generalisation of Lemma 9 appearing on page 217 of +[J.W.S. Cassels, *Some metrical theorems in Diophantine approximation. I*](cassels1950). + +NB: The `set : α` type ascription is present because of issue #16932 on GitHub. -/ +theorem blimsup_cthickening_mul_ae_eq + (p : ℕ → Prop) (s : ℕ → set α) {M : ℝ} (hM : 0 < M) (r : ℕ → ℝ) (hr : tendsto r at_top (𝓝 0)) : + (blimsup (λ i, cthickening (M * r i) (s i)) at_top p : set α) =ᵐ[μ] + (blimsup (λ i, cthickening (r i) (s i)) at_top p : set α) := +begin + have : ∀ (p : ℕ → Prop) {r : ℕ → ℝ} (hr : tendsto r at_top (𝓝[>] 0)), + (blimsup (λ i, cthickening (M * r i) (s i)) at_top p : set α) =ᵐ[μ] + (blimsup (λ i, cthickening (r i) (s i)) at_top p : set α), + { clear p hr r, intros p r hr, + have hr' : tendsto (λ i, M * r i) at_top (𝓝[>] 0), + { convert tendsto_nhds_within_Ioi.const_mul hM hr; simp only [mul_zero], }, + refine eventually_le_antisymm_iff.mpr ⟨_, _⟩, + { exact blimsup_cthickening_ae_le_of_eventually_mul_le μ p (inv_pos.mpr hM) hr' + (eventually_of_forall $ λ i, by rw inv_mul_cancel_left₀ hM.ne' (r i)), }, + { exact blimsup_cthickening_ae_le_of_eventually_mul_le μ p hM hr + (eventually_of_forall $ λ i, le_refl _), }, }, + let r' : ℕ → ℝ := λ i, if 0 < r i then r i else 1/((i : ℝ) + 1), + have hr' : tendsto r' at_top (𝓝[>] 0), + { refine tendsto_nhds_within_iff.mpr ⟨tendsto.if' hr tendsto_one_div_add_at_top_nhds_0_nat, + eventually_of_forall $ λ i, _⟩, + by_cases hi : 0 < r i, + { simp [hi, r'], }, + { simp only [hi, r', one_div, mem_Ioi, if_false, inv_pos], positivity, }, }, + have h₀ : ∀ i, (p i ∧ 0 < r i) → cthickening (r i) (s i) = cthickening (r' i) (s i), + { rintros i ⟨-, hi⟩, congr, change r i = ite (0 < r i) (r i) _, simp [hi], }, + have h₁ : ∀ i, (p i ∧ 0 < r i) → cthickening (M * r i) (s i) = cthickening (M * r' i) (s i), + { rintros i ⟨-, hi⟩, simp only [hi, mul_ite, if_true], }, + have h₂ : ∀ i, (p i ∧ r i ≤ 0) → cthickening (M * r i) (s i) = cthickening (r i) (s i), + { rintros i ⟨-, hi⟩, + have hi' : M * r i ≤ 0 := mul_nonpos_of_nonneg_of_nonpos hM.le hi, + rw [cthickening_of_nonpos hi, cthickening_of_nonpos hi'], }, + have hp : p = λ i, (p i ∧ 0 < r i) ∨ (p i ∧ r i ≤ 0), + { ext i, simp [← and_or_distrib_left, lt_or_le 0 (r i)], }, + rw [hp, blimsup_or_eq_sup, blimsup_or_eq_sup, sup_eq_union, + blimsup_congr (eventually_of_forall h₀), blimsup_congr (eventually_of_forall h₁), + blimsup_congr (eventually_of_forall h₂)], + exact ae_eq_set_union (this (λ i, p i ∧ 0 < r i) hr') (ae_eq_refl _), +end diff --git a/src/measure_theory/measure/doubling.lean b/src/measure_theory/measure/doubling.lean index b9562c46077a0..c293a057d4dd6 100644 --- a/src/measure_theory/measure/doubling.lean +++ b/src/measure_theory/measure/doubling.lean @@ -25,7 +25,7 @@ This file records basic files on doubling measures. noncomputable theory open set filter metric measure_theory topological_space -open_locale nnreal topological_space +open_locale ennreal nnreal topological_space /-- A measure `μ` is said to be a doubling measure if there exists a constant `C` such that for all sufficiently small radii `ε`, and for any centre, the measure of a ball of radius `2 * ε` is @@ -87,6 +87,9 @@ values other than `2`. -/ def scaling_constant_of (K : ℝ) : ℝ≥0 := max (classical.some $ exists_eventually_forall_measure_closed_ball_le_mul μ K) 1 +@[simp] lemma one_le_scaling_constant_of (K : ℝ) : 1 ≤ scaling_constant_of μ K := +le_max_of_le_right $ le_refl 1 + lemma eventually_measure_mul_le_scaling_constant_of_mul (K : ℝ) : ∃ (R : ℝ), 0 < R ∧ ∀ x t r (ht : t ∈ Ioc 0 K) (hr : r ≤ R), μ (closed_ball x (t * r)) ≤ scaling_constant_of μ K * μ (closed_ball x r) := @@ -104,6 +107,24 @@ begin exact ennreal.mul_le_mul (ennreal.coe_le_coe.2 (le_max_left _ _)) le_rfl } end +lemma eventually_measure_le_scaling_constant_mul (K : ℝ) : + ∀ᶠ r in 𝓝[>] 0, ∀ x, + μ (closed_ball x (K * r)) ≤ scaling_constant_of μ K * μ (closed_ball x r) := +begin + filter_upwards [classical.some_spec (exists_eventually_forall_measure_closed_ball_le_mul μ K)] + with r hr x, + exact (hr x K le_rfl).trans (ennreal.mul_le_mul (ennreal.coe_le_coe.2 (le_max_left _ _)) le_rfl) +end + +lemma eventually_measure_le_scaling_constant_mul' (K : ℝ) (hK : 0 < K) : + ∀ᶠ r in 𝓝[>] 0, ∀ x, + μ (closed_ball x r) ≤ scaling_constant_of μ K⁻¹ * μ (closed_ball x (K * r)) := +begin + convert eventually_nhds_within_pos_mul_left hK (eventually_measure_le_scaling_constant_mul μ K⁻¹), + ext, + simp [inv_mul_cancel_left₀ hK.ne'], +end + /-- A scale below which the doubling measure `μ` satisfies good rescaling properties when one multiplies the radius of balls by at most `K`, as stated in `measure_mul_le_scaling_constant_of_mul`. -/ diff --git a/src/measure_theory/measure/measure_space_def.lean b/src/measure_theory/measure/measure_space_def.lean index 816b05437d4f7..1da5a90cb851b 100644 --- a/src/measure_theory/measure/measure_space_def.lean +++ b/src/measure_theory/measure/measure_space_def.lean @@ -398,6 +398,10 @@ lemma ae_eq_set_inter {s' t' : set α} (h : s =ᵐ[μ] t) (h' : s' =ᵐ[μ] t') (s ∩ s' : set α) =ᵐ[μ] (t ∩ t' : set α) := h.inter h' +lemma ae_eq_set_union {s' t' : set α} (h : s =ᵐ[μ] t) (h' : s' =ᵐ[μ] t') : + (s ∪ s' : set α) =ᵐ[μ] (t ∪ t' : set α) := +h.union h' + @[to_additive] lemma _root_.set.mul_indicator_ae_eq_one {M : Type*} [has_one M] {f : α → M} {s : set α} (h : s.mul_indicator f =ᵐ[μ] 1) : μ (s ∩ function.mul_support f) = 0 := diff --git a/src/order/filter/basic.lean b/src/order/filter/basic.lean index 0577d370940b8..607653b6c458b 100644 --- a/src/order/filter/basic.lean +++ b/src/order/filter/basic.lean @@ -2575,6 +2575,15 @@ begin exacts [hp₀ h, hp₁ h], end +lemma tendsto.if' {α β : Type*} {l₁ : filter α} {l₂ : filter β} {f g : α → β} {p : α → Prop} + [decidable_pred p] (hf : tendsto f l₁ l₂) (hg : tendsto g l₁ l₂) : + tendsto (λ a, if p a then f a else g a) l₁ l₂ := +begin + replace hf : tendsto f (l₁ ⊓ 𝓟 {x | p x}) l₂ := tendsto_inf_left hf, + replace hg : tendsto g (l₁ ⊓ 𝓟 {x | ¬ p x}) l₂ := tendsto_inf_left hg, + exact hf.if hg, +end + lemma tendsto.piecewise {l₁ : filter α} {l₂ : filter β} {f g : α → β} {s : set α} [∀ x, decidable (x ∈ s)] (h₀ : tendsto f (l₁ ⊓ 𝓟 s) l₂) (h₁ : tendsto g (l₁ ⊓ 𝓟 sᶜ) l₂) : diff --git a/src/order/liminf_limsup.lean b/src/order/liminf_limsup.lean index ff6dd40b3ec18..d0200be270336 100644 --- a/src/order/liminf_limsup.lean +++ b/src/order/liminf_limsup.lean @@ -439,6 +439,20 @@ begin exact eventually_congr (h.mono $ λ x hx, by simp [hx]) end +lemma blimsup_congr {f : filter β} {u v : β → α} {p : β → Prop} (h : ∀ᶠ a in f, p a → u a = v a) : + blimsup u f p = blimsup v f p := +begin + rw blimsup_eq, + congr' with b, + refine eventually_congr (h.mono $ λ x hx, ⟨λ h₁ h₂, _, λ h₁ h₂, _⟩), + { rw ← hx h₂, exact h₁ h₂, }, + { rw hx h₂, exact h₁ h₂, }, +end + +lemma bliminf_congr {f : filter β} {u v : β → α} {p : β → Prop} (h : ∀ᶠ a in f, p a → u a = v a) : + bliminf u f p = bliminf v f p := +@blimsup_congr αᵒᵈ _ _ _ _ _ _ h + lemma liminf_congr {α : Type*} [conditionally_complete_lattice β] {f : filter α} {u v : α → β} (h : ∀ᶠ a in f, u a = v a) : liminf u f = liminf v f := @limsup_congr βᵒᵈ _ _ _ _ _ h @@ -544,6 +558,12 @@ begin exact (le_infi_iff.mp (ha s) hs).trans (by simpa only [supr₂_le_iff, and_imp]), }, end +lemma blimsup_eq_infi_bsupr_of_nat {p : ℕ → Prop} {u : ℕ → α} : + blimsup u at_top p = ⨅ i, ⨆ j (hj : p j ∧ i ≤ j), u j := +by simp only [blimsup_eq_limsup_subtype, mem_preimage, mem_Ici, function.comp_app, cinfi_pos, + supr_subtype, (at_top_basis.comap (coe : {x | p x} → ℕ)).limsup_eq_infi_supr, mem_set_of_eq, + subtype.coe_mk, supr_and] + /-- 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 : β → α} : liminf u f = ⨆ s ∈ f, ⨅ a ∈ s, u a := @@ -563,6 +583,10 @@ lemma bliminf_eq_supr_binfi {f : filter β} {p : β → Prop} {u : β → α} : bliminf u f p = ⨆ s ∈ f, ⨅ b (hb : p b ∧ b ∈ s), u b := @blimsup_eq_infi_bsupr αᵒᵈ β _ f p u +lemma bliminf_eq_supr_binfi_of_nat {p : ℕ → Prop} {u : ℕ → α} : + bliminf u at_top p = ⨆ i, ⨅ j (hj : p j ∧ i ≤ j), u j := +@blimsup_eq_infi_bsupr_of_nat αᵒᵈ _ p u + lemma limsup_eq_Inf_Sup {ι R : Type*} (F : filter ι) [complete_lattice R] (a : ι → R) : limsup a F = Inf ((λ I, Sup (a '' I)) '' F.sets) := begin @@ -800,6 +824,28 @@ lemma cofinite.liminf_set_eq : liminf s cofinite = { x | { n | x ∉ s n }.finite } := by simp only [← cofinite.bliminf_true s, cofinite.bliminf_set_eq, true_and] +lemma exists_forall_mem_of_has_basis_mem_blimsup + {l : filter β} {b : ι → set β} {q : ι → Prop} (hl : l.has_basis q b) + {u : β → set α} {p : β → Prop} {x : α} (hx : x ∈ blimsup u l p) : + ∃ f : {i | q i} → β, ∀ i, x ∈ u (f i) ∧ p (f i) ∧ f i ∈ b i := +begin + rw blimsup_eq_infi_bsupr at hx, + simp only [supr_eq_Union, infi_eq_Inter, mem_Inter, mem_Union, exists_prop] at hx, + choose g hg hg' using hx, + refine ⟨λ (i : {i | q i}), g (b i) (hl.mem_of_mem i.2), λ i, ⟨_, _⟩⟩, + { exact hg' (b i) (hl.mem_of_mem i.2), }, + { exact hg (b i) (hl.mem_of_mem i.2), }, +end + +lemma exists_forall_mem_of_has_basis_mem_blimsup' + {l : filter β} {b : ι → set β} (hl : l.has_basis (λ _, true) b) + {u : β → set α} {p : β → Prop} {x : α} (hx : x ∈ blimsup u l p) : + ∃ f : ι → β, ∀ i, x ∈ u (f i) ∧ p (f i) ∧ f i ∈ b i := +begin + obtain ⟨f, hf⟩ := exists_forall_mem_of_has_basis_mem_blimsup hl hx, + exact ⟨λ i, f ⟨i, trivial⟩, λ i, hf ⟨i, trivial⟩⟩, +end + end set_lattice section conditionally_complete_linear_order diff --git a/src/topology/metric_space/hausdorff_distance.lean b/src/topology/metric_space/hausdorff_distance.lean index 07e80930c5727..1643b58e357af 100644 --- a/src/topology/metric_space/hausdorff_distance.lean +++ b/src/topology/metric_space/hausdorff_distance.lean @@ -1294,8 +1294,19 @@ begin simpa using hx, end +lemma cthickening_subset_Union_closed_ball_of_lt + {α : Type*} [pseudo_metric_space α] (E : set α) {δ δ' : ℝ} (hδ₀ : 0 < δ') (hδδ' : δ < δ') : + cthickening δ E ⊆ ⋃ x ∈ E, closed_ball x δ' := +begin + refine (cthickening_subset_thickening' hδ₀ hδδ' E).trans (λ x hx, _), + obtain ⟨y, hy₁, hy₂⟩ := mem_thickening_iff.mp hx, + exact mem_Union₂.mpr ⟨y, hy₁, hy₂.le⟩, +end + /-- The closed thickening of a compact set `E` is the union of the balls `closed_ball x δ` over -`x ∈ E`. -/ +`x ∈ E`. + +See also `metric.cthickening_eq_bUnion_closed_ball`. -/ lemma _root_.is_compact.cthickening_eq_bUnion_closed_ball {α : Type*} [pseudo_metric_space α] {δ : ℝ} {E : set α} (hE : is_compact E) (hδ : 0 ≤ δ) : cthickening δ E = ⋃ x ∈ E, closed_ball x δ := @@ -1312,6 +1323,26 @@ begin exact mem_bUnion yE D2, end +lemma cthickening_eq_bUnion_closed_ball + {α : Type*} [pseudo_metric_space α] [proper_space α] (E : set α) (hδ : 0 ≤ δ) : + cthickening δ E = ⋃ x ∈ closure E, closed_ball x δ := +begin + rcases eq_empty_or_nonempty E with rfl|hne, + { simp only [cthickening_empty, Union_false, Union_empty, closure_empty], }, + rw ← cthickening_closure, + refine subset.antisymm (λ x hx, _) (Union₂_subset $ λ x hx, closed_ball_subset_cthickening hx _), + obtain ⟨y, yE, hy⟩ : ∃ y ∈ closure E, inf_dist x (closure E) = dist x y := + is_closed_closure.exists_inf_dist_eq_dist (closure_nonempty_iff.mpr hne) x, + replace hy : dist x y ≤ δ := (ennreal.of_real_le_of_real_iff hδ).mp + (((congr_arg ennreal.of_real hy.symm).le.trans ennreal.of_real_to_real_le).trans hx), + exact mem_bUnion yE hy, +end + +lemma _root_.is_closed.cthickening_eq_bUnion_closed_ball + {α : Type*} [pseudo_metric_space α] [proper_space α] {E : set α} (hE : is_closed E) (hδ : 0 ≤ δ) : + cthickening δ E = ⋃ x ∈ E, closed_ball x δ := +by rw [cthickening_eq_bUnion_closed_ball E hδ, hE.closure_eq] + /-- For the equality, see `inf_edist_cthickening`. -/ lemma inf_edist_le_inf_edist_cthickening_add : inf_edist x s ≤ inf_edist x (cthickening δ s) + ennreal.of_real δ := diff --git a/src/topology/order/basic.lean b/src/topology/order/basic.lean index 5d09aabc6822f..460b8dd82d41c 100644 --- a/src/topology/order/basic.lean +++ b/src/topology/order/basic.lean @@ -583,6 +583,41 @@ lemma filter.tendsto.min {b : filter β} {a₁ a₂ : α} (hf : tendsto f b ( tendsto (λb, min (f b) (g b)) b (𝓝 (min a₁ a₂)) := (continuous_min.tendsto (a₁, a₂)).comp (hf.prod_mk_nhds hg) +lemma filter.tendsto.max_right {l : filter β} {a : α} (h : tendsto f l (𝓝 a)) : + tendsto (λ i, max a (f i)) l (𝓝 a) := +by { convert ((continuous_max.comp (@continuous.prod.mk α α _ _ a)).tendsto a).comp h, simp, } + +lemma filter.tendsto.max_left {l : filter β} {a : α} (h : tendsto f l (𝓝 a)) : + tendsto (λ i, max (f i) a) l (𝓝 a) := +by { simp_rw max_comm _ a, exact h.max_right, } + +lemma filter.tendsto_nhds_max_right {l : filter β} {a : α} (h : tendsto f l (𝓝[>] a)) : + tendsto (λ i, max a (f i)) l (𝓝[>] a) := +begin + obtain ⟨h₁ : tendsto f l (𝓝 a), h₂ : ∀ᶠ i in l, f i ∈ Ioi a⟩ := tendsto_nhds_within_iff.mp h, + exact tendsto_nhds_within_iff.mpr ⟨h₁.max_right, h₂.mono $ λ i hi, lt_max_of_lt_right hi⟩, +end + +lemma filter.tendsto_nhds_max_left {l : filter β} {a : α} (h : tendsto f l (𝓝[>] a)) : + tendsto (λ i, max (f i) a) l (𝓝[>] a) := +by { simp_rw max_comm _ a, exact filter.tendsto_nhds_max_right h, } + +lemma filter.tendsto.min_right {l : filter β} {a : α} (h : tendsto f l (𝓝 a)) : + tendsto (λ i, min a (f i)) l (𝓝 a) := +@filter.tendsto.max_right αᵒᵈ β _ _ _ f l a h + +lemma filter.tendsto.min_left {l : filter β} {a : α} (h : tendsto f l (𝓝 a)) : + tendsto (λ i, min (f i) a) l (𝓝 a) := +@filter.tendsto.max_left αᵒᵈ β _ _ _ f l a h + +lemma filter.tendsto_nhds_min_right {l : filter β} {a : α} (h : tendsto f l (𝓝[<] a)) : + tendsto (λ i, min a (f i)) l (𝓝[<] a) := +@filter.tendsto_nhds_max_right αᵒᵈ β _ _ _ f l a h + +lemma filter.tendsto_nhds_min_left {l : filter β} {a : α} (h : tendsto f l (𝓝[<] a)) : + tendsto (λ i, min (f i) a) l (𝓝[<] a) := +@filter.tendsto_nhds_max_left αᵒᵈ β _ _ _ f l a h + lemma dense.exists_lt [no_min_order α] {s : set α} (hs : dense s) (x : α) : ∃ y ∈ s, y < x := hs.exists_mem_open is_open_Iio (exists_lt x) From ae3588fa1541923f6798f21a577c61ad61524947 Mon Sep 17 00:00:00 2001 From: Oliver Nash Date: Wed, 14 Dec 2022 17:11:38 +0000 Subject: [PATCH 0044/1291] feat(dynamics/ergodic/add_circle): generalise `add_circle.ae_empty_or_univ_of_forall_vadd_eq_self` slightly (#17907) I need this generalised version for an application. --- src/dynamics/ergodic/add_circle.lean | 14 ++++---- src/measure_theory/group/action.lean | 33 ++++++++++++++++++- src/measure_theory/group/add_circle.lean | 8 ++--- src/measure_theory/measure/measure_space.lean | 31 +++++++++++++++++ 4 files changed, 74 insertions(+), 12 deletions(-) diff --git a/src/dynamics/ergodic/add_circle.lean b/src/dynamics/ergodic/add_circle.lean index 389be1df41d5a..5a22777975888 100644 --- a/src/dynamics/ergodic/add_circle.lean +++ b/src/dynamics/ergodic/add_circle.lean @@ -34,12 +34,13 @@ namespace add_circle variables {T : ℝ} [hT : fact (0 < T)] include hT -/-- If a null-measurable subset of the circle is invariant under rotation by a family of rational -angles with denominators tending to infinity, then it must be almost empty or almost full. -/ -lemma ae_empty_or_univ_of_forall_vadd_eq_self +/-- If a null-measurable subset of the circle is almost invariant under rotation by a family of +rational angles with denominators tending to infinity, then it must be almost empty or almost full. +-/ +lemma ae_empty_or_univ_of_forall_vadd_ae_eq_self {s : set $ add_circle T} (hs : null_measurable_set s volume) {ι : Type*} {l : filter ι} [l.ne_bot] {u : ι → add_circle T} - (hu₁ : ∀ i, (u i) +ᵥ s = s) (hu₂ : tendsto (add_order_of ∘ u) l at_top) : + (hu₁ : ∀ i, ((u i) +ᵥ s : set _) =ᵐ[volume] s) (hu₂ : tendsto (add_order_of ∘ u) l at_top) : s =ᵐ[volume] (∅ : set $ add_circle T) ∨ s =ᵐ[volume] univ := begin /- Sketch of proof: @@ -102,10 +103,11 @@ lemma ergodic_zsmul {n : ℤ} (hn : 1 < |n|) : ergodic (λ (y : add_circle T), n { exact λ j, add_order_of_div_of_gcd_eq_one (pow_pos (pos_of_gt hn) j) (gcd_one_left _), }, have hnu : ∀ j, n^j • (u j) = 0 := λ j, by rw [← add_order_of_dvd_iff_zsmul_eq_zero, hu₀, int.coe_nat_pow, ← int.abs_eq_nat_abs, ← abs_pow, abs_dvd], - have hu₁ : ∀ j, (u j) +ᵥ s = s := λ j, vadd_eq_self_of_preimage_zsmul_eq_self hs' (hnu j), + have hu₁ : ∀ j, ((u j) +ᵥ s : set _) =ᵐ[volume] s := + λ j, by rw vadd_eq_self_of_preimage_zsmul_eq_self hs' (hnu j), have hu₂ : tendsto (λ j, add_order_of $ u j) at_top at_top, { simp_rw hu₀, exact nat.tendsto_pow_at_top_at_top_of_one_lt hn, }, - exact ae_empty_or_univ_of_forall_vadd_eq_self hs.null_measurable_set hu₁ hu₂, + exact ae_empty_or_univ_of_forall_vadd_ae_eq_self hs.null_measurable_set hu₁ hu₂, end, .. measure_preserving_zsmul volume (abs_pos.mp $ lt_trans zero_lt_one hn), } diff --git a/src/measure_theory/group/action.lean b/src/measure_theory/group/action.lean index 746e81e3a47a8..a0c718bdc1c63 100644 --- a/src/measure_theory/group/action.lean +++ b/src/measure_theory/group/action.lean @@ -22,7 +22,7 @@ open measure_theory measure_theory.measure set function namespace measure_theory -variables {G M α : Type*} +variables {G M α : Type*} {s : set α} /-- A measure `μ : measure α` is invariant under an additive action of `M` on `α` if for any measurable set `s : set α` and `c : M`, the measure of its preimage under `λ x, c +ᵥ x` is equal to @@ -193,4 +193,35 @@ by rw [← not_iff_not, ← ne.def, ← pos_iff_ne_zero, end is_minimal +lemma smul_ae_eq_self_of_mem_zpowers + {x y : G} (hs : (x • s : set α) =ᵐ[μ] s) (hy : y ∈ subgroup.zpowers x) : + (y • s : set α) =ᵐ[μ] s := +begin + obtain ⟨k, rfl⟩ := subgroup.mem_zpowers_iff.mp hy, + let e : α ≃ α := mul_action.to_perm_hom G α x, + have he : quasi_measure_preserving e μ μ := + (measure_preserving_smul x μ).quasi_measure_preserving, + have he' : quasi_measure_preserving e.symm μ μ := + (measure_preserving_smul x⁻¹ μ).quasi_measure_preserving, + simpa only [mul_action.to_perm_hom_apply, mul_action.to_perm_apply, image_smul, + ← monoid_hom.map_zpow] using he.image_zpow_ae_eq he' k hs, +end + +lemma vadd_ae_eq_self_of_mem_zmultiples {G : Type*} [measurable_space G] + [add_group G] [add_action G α] [vadd_invariant_measure G α μ] [has_measurable_vadd G α] + {x y : G} (hs : (x +ᵥ s : set α) =ᵐ[μ] s) (hy : y ∈ add_subgroup.zmultiples x) : + (y +ᵥ s : set α) =ᵐ[μ] s := +begin + letI : measurable_space (multiplicative G) := (by apply_instance : measurable_space G), + letI : smul_invariant_measure (multiplicative G) α μ := + ⟨λ g, vadd_invariant_measure.measure_preimage_vadd μ (multiplicative.to_add g)⟩, + letI : has_measurable_smul (multiplicative G) α := + { measurable_const_smul := λ g, measurable_const_vadd (multiplicative.to_add g), + measurable_smul_const := λ a, @measurable_vadd_const (multiplicative G) α + (by apply_instance : has_vadd G α) _ _ (by apply_instance : has_measurable_vadd G α) a }, + exact @smul_ae_eq_self_of_mem_zpowers (multiplicative G) α _ _ _ _ _ _ _ _ _ _ hs hy, +end + +attribute [to_additive vadd_ae_eq_self_of_mem_zmultiples] smul_ae_eq_self_of_mem_zpowers + end measure_theory diff --git a/src/measure_theory/group/add_circle.lean b/src/measure_theory/group/add_circle.lean index 291a425d2728d..93388f293ccf4 100644 --- a/src/measure_theory/group/add_circle.lean +++ b/src/measure_theory/group/add_circle.lean @@ -93,17 +93,15 @@ begin exact le_refl _, }, end -lemma volume_of_add_preimage_eq (s I : set $ add_circle T) - (u x : add_circle T) (hu : is_of_fin_add_order u) (hs : u +ᵥ s = s) +lemma volume_of_add_preimage_eq (s I : set $ add_circle T) (u x : add_circle T) + (hu : is_of_fin_add_order u) (hs : (u +ᵥ s : set $ add_circle T) =ᵐ[volume] s) (hI : I =ᵐ[volume] ball x (T / (2 * add_order_of u))) : volume s = add_order_of u • volume (s ∩ I) := begin let G := add_subgroup.zmultiples u, haveI : fintype G := @fintype.of_finite _ hu.finite_zmultiples, have hsG : ∀ (g : G), (g +ᵥ s : set $ add_circle T) =ᵐ[volume] s, - { rintros ⟨y, hy⟩, - convert ae_eq_refl s, - exact vadd_eq_self_of_mem_zmultiples hy hs, }, + { rintros ⟨y, hy⟩, exact (vadd_ae_eq_self_of_mem_zmultiples hs hy : _), }, rw [(is_add_fundamental_domain_of_ae_ball I u x hu hI).measure_eq_card_smul_of_vadd_ae_eq_self s hsG, add_order_eq_card_zmultiples' u, nat.card_eq_fintype_card], end diff --git a/src/measure_theory/measure/measure_space.lean b/src/measure_theory/measure/measure_space.lean index 82ab393b11391..3401782f58569 100644 --- a/src/measure_theory/measure/measure_space.lean +++ b/src/measure_theory/measure/measure_space.lean @@ -2126,6 +2126,37 @@ lemma preimage_null (h : quasi_measure_preserving f μa μb) {s : set β} (hs : μa (f ⁻¹' s) = 0 := preimage_null_of_map_null h.ae_measurable (h.2 hs) +lemma preimage_mono_ae {s t : set β} (hf : quasi_measure_preserving f μa μb) (h : s ≤ᵐ[μb] t) : + f⁻¹' s ≤ᵐ[μa] f⁻¹' t := +eventually_map.mp $ eventually.filter_mono (tendsto_ae_map hf.ae_measurable) + (eventually.filter_mono hf.ae_map_le h) + +lemma preimage_ae_eq {s t : set β} (hf : quasi_measure_preserving f μa μb) (h : s =ᵐ[μb] t) : + f⁻¹' s =ᵐ[μa] f⁻¹' t := +eventually_le.antisymm (hf.preimage_mono_ae h.le) (hf.preimage_mono_ae h.symm.le) + +lemma preimage_iterate_ae_eq {s : set α} {f : α → α} (hf : quasi_measure_preserving f μ μ) (k : ℕ) + (hs : f⁻¹' s =ᵐ[μ] s) : (f^[k])⁻¹' s =ᵐ[μ] s := +begin + induction k with k ih, { simp, }, + rw [iterate_succ, preimage_comp], + exact eventually_eq.trans (hf.preimage_ae_eq ih) hs, +end + +lemma image_zpow_ae_eq {s : set α} {e : α ≃ α} (he : quasi_measure_preserving e μ μ) + (he' : quasi_measure_preserving e.symm μ μ) (k : ℤ) (hs : e '' s =ᵐ[μ] s) : ⇑(e^k) '' s =ᵐ[μ] s := +begin + rw equiv.image_eq_preimage, + obtain ⟨k, rfl | rfl⟩ := k.eq_coe_or_neg, + { replace hs : ⇑(e⁻¹)⁻¹' s =ᵐ[μ] s, by rwa equiv.image_eq_preimage at hs, + replace he' : (⇑(e⁻¹)^[k])⁻¹' s =ᵐ[μ] s := he'.preimage_iterate_ae_eq k hs, + rwa [equiv.perm.iterate_eq_pow e⁻¹ k, inv_pow e k] at he', }, + { rw [zpow_neg, zpow_coe_nat], + replace hs : e⁻¹' s =ᵐ[μ] s, { convert he.preimage_ae_eq hs.symm, rw equiv.preimage_image, }, + replace he : (⇑e^[k])⁻¹' s =ᵐ[μ] s := he.preimage_iterate_ae_eq k hs, + rwa [equiv.perm.iterate_eq_pow e k] at he, }, +end + lemma limsup_preimage_iterate_ae_eq {f : α → α} (hf : quasi_measure_preserving f μ μ) (hs : f⁻¹' s =ᵐ[μ] s) : -- Need `@` below because of diamond; see gh issue #16932 From 7d6cd0411fa5334c211f4657e971f2666d08b15a Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Wed, 14 Dec 2022 17:12:36 +0000 Subject: [PATCH 0045/1291] feat(ring_theory/power_series): `X s` commutes with any power series (#17935) Use this fact to golf a proof and generalize `order_eq_multiplicity_X` from `[comm_semiring R]` to `[semiring R]`. Should we redefine `order` to be `multiplicity` in all cases? --- src/ring_theory/power_series/basic.lean | 26 ++++++++++++++++++------- 1 file changed, 19 insertions(+), 7 deletions(-) diff --git a/src/ring_theory/power_series/basic.lean b/src/ring_theory/power_series/basic.lean index e512ad9a7cc90..5752dc54e58b4 100644 --- a/src/ring_theory/power_series/basic.lean +++ b/src/ring_theory/power_series/basic.lean @@ -209,6 +209,16 @@ begin exact le_add_left le_rfl end +@[simp] lemma commute_monomial {a : R} {n} : + commute φ (monomial R n a) ↔ ∀ m, commute (coeff R m φ) a := +begin + refine ext_iff.trans ⟨λ h m, _, λ h m, _⟩, + { have := h (m + n), + rwa [coeff_add_mul_monomial, add_comm, coeff_add_monomial_mul] at this }, + { rw [coeff_mul_monomial, coeff_monomial_mul], + split_ifs; [apply h, refl] } +end + protected lemma one_mul : (1 : mv_power_series σ R) * φ = φ := ext $ λ n, by simpa using coeff_add_monomial_mul 0 n φ 1 @@ -323,6 +333,9 @@ coeff_monomial_same _ _ lemma coeff_zero_X (s : σ) : coeff R (0 : σ →₀ ℕ) (X s : mv_power_series σ R) = 0 := by { rw [coeff_X, if_neg], intro h, exact one_ne_zero (single_eq_zero.mp h.symm) } +lemma commute_X (φ : mv_power_series σ R) (s : σ) : commute φ (X s) := +φ.commute_monomial.mpr $ λ m, commute.one_right _ + lemma X_def (s : σ) : X s = monomial R (single s 1) 1 := rfl lemma X_pow_eq (s : σ) (n : ℕ) : @@ -353,11 +366,8 @@ begin end lemma coeff_zero_X_mul (φ : mv_power_series σ R) (s : σ) : - coeff R (0 : σ →₀ ℕ) (X s * φ) = 0 := -begin - have : ¬single s 1 ≤ 0, from λ h, by simpa using h s, - simp only [X, coeff_monomial_mul, if_neg this] -end + coeff R (0 : σ →₀ ℕ) (X s * φ) = 0 := +by rw [← (φ.commute_X s).eq, coeff_zero_mul_X] variables (σ) (R) @@ -1010,6 +1020,8 @@ variable {R} /-- The variable of the formal power series ring.-/ def X : power_series R := mv_power_series.X () +lemma commute_X (φ : power_series R) : commute φ X := φ.commute_X _ + @[simp] lemma coeff_zero_eq_constant_coeff : ⇑(coeff R 0) = constant_coeff R := by { rw [coeff, finsupp.single_zero], refl } @@ -1858,7 +1870,7 @@ begin simpa [part_enat.coe_lt_iff] using λ _, hn } end -lemma order_eq_multiplicity_X {R : Type*} [comm_semiring R] (φ : power_series R) : +lemma order_eq_multiplicity_X {R : Type*} [semiring R] (φ : power_series R) : order φ = multiplicity X φ := begin rcases eq_or_ne φ 0 with rfl|hφ, @@ -1872,7 +1884,7 @@ begin (order_finite_iff_ne_zero.mpr hφ)) (part_enat.find_le _ _ _), rintro ⟨ψ, H⟩, have := congr_arg (coeff R n) H, - rw [mul_comm, coeff_mul_of_lt_order, ←hn] at this, + rw [← (ψ.commute_X.pow_right _).eq, coeff_mul_of_lt_order, ←hn] at this, { exact coeff_order _ this }, { rw [X_pow_eq, order_monomial], split_ifs, From f2c46488c2870f03a49584622b0bee2e95149a3e Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 14 Dec 2022 17:12:38 +0000 Subject: [PATCH 0046/1291] =?UTF-8?q?feat(order/succ=5Fpred/basic):=20`a?= =?UTF-8?q?=20=E2=A9=BF=20b=20=E2=86=92=20b=20=E2=89=A4=20succ=20a`=20(#17?= =?UTF-8?q?940)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit If `a` weakly covers `b`, then `b` is at most the successor of `a`. --- src/order/cover.lean | 5 +++++ src/order/succ_pred/basic.lean | 14 ++++++++++++++ 2 files changed, 19 insertions(+) diff --git a/src/order/cover.lean b/src/order/cover.lean index b8b9b6e84213f..5a9f6fa0729b4 100644 --- a/src/order/cover.lean +++ b/src/order/cover.lean @@ -272,6 +272,11 @@ lemma covby_iff_wcovby_and_ne : a ⋖ b ↔ a ⩿ b ∧ a ≠ b := lemma wcovby_iff_covby_or_eq : a ⩿ b ↔ a ⋖ b ∨ a = b := by rw [le_antisymm_iff, wcovby_iff_covby_or_le_and_le] +lemma wcovby_iff_eq_or_covby : a ⩿ b ↔ a = b ∨ a ⋖ b := wcovby_iff_covby_or_eq.trans or.comm + +alias wcovby_iff_covby_or_eq ↔ wcovby.covby_or_eq _ +alias wcovby_iff_eq_or_covby ↔ wcovby.eq_or_covby _ + lemma covby.eq_or_eq (h : a ⋖ b) (h2 : a ≤ c) (h3 : c ≤ b) : c = a ∨ c = b := h.wcovby.eq_or_eq h2 h3 diff --git a/src/order/succ_pred/basic.lean b/src/order/succ_pred/basic.lean index c9f22fb29d788..eff74e663ee47 100644 --- a/src/order/succ_pred/basic.lean +++ b/src/order/succ_pred/basic.lean @@ -314,6 +314,13 @@ end lemma _root_.covby.succ_eq (h : a ⋖ b) : succ a = b := (succ_le_of_lt h.lt).eq_of_not_lt $ λ h', h.2 (lt_succ_of_not_is_max h.lt.not_is_max) h' +lemma _root_.wcovby.le_succ (h : a ⩿ b) : b ≤ succ a := +begin + obtain h | rfl := h.covby_or_eq, + { exact h.succ_eq.ge }, + { exact le_succ _ } +end + lemma le_succ_iff_eq_or_le : a ≤ succ b ↔ a = succ b ∨ a ≤ b := begin by_cases hb : is_max b, @@ -550,6 +557,13 @@ end lemma _root_.covby.pred_eq {a b : α} (h : a ⋖ b) : pred b = a := (le_pred_of_lt h.lt).eq_of_not_gt $ λ h', h.2 h' $ pred_lt_of_not_is_min h.lt.not_is_min +lemma _root_.wcovby.pred_le (h : a ⩿ b) : pred b ≤ a := +begin + obtain h | rfl := h.covby_or_eq, + { exact h.pred_eq.le }, + { exact pred_le _ } +end + lemma pred_le_iff_eq_or_le : pred a ≤ b ↔ b = pred a ∨ a ≤ b := begin by_cases ha : is_min a, From aba57d4d3dae35460225919dcd82fe91355162f9 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 14 Dec 2022 20:14:27 +0000 Subject: [PATCH 0047/1291] feat(combinatorics/double_counting): Special cases (#17647) GIve special cases of `card_mul_le_card_mul` and `card_mul_le_card_mul'` when `m = n = 1`. --- src/combinatorics/double_counting.lean | 40 ++++++++++++++++++++++++-- 1 file changed, 38 insertions(+), 2 deletions(-) diff --git a/src/combinatorics/double_counting.lean b/src/combinatorics/double_counting.lean index 99e30deda21c4..3c042ba9a7ed4 100644 --- a/src/combinatorics/double_counting.lean +++ b/src/combinatorics/double_counting.lean @@ -26,14 +26,16 @@ and `t`. * `card_mul_eq_card_mul`: Equality combination of the previous. -/ -open finset function +open finset function relator open_locale big_operators +variables {α β : Type*} + /-! ### Bipartite graph -/ namespace finset section bipartite -variables {α β : Type*} (r : α → β → Prop) (s : finset α) (t : finset β) (a a' : α) (b b' : β) +variables (r : α → β → Prop) (s : finset α) (t : finset β) (a a' : α) (b b' : β) [decidable_pred (r a)] [Π a, decidable (r a b)] {m n : ℕ} /-- Elements of `s` which are "below" `b` according to relation `r`. -/ @@ -45,6 +47,12 @@ def bipartite_above : finset β := t.filter (r a) lemma bipartite_below_swap : t.bipartite_below (swap r) a = t.bipartite_above r a := rfl lemma bipartite_above_swap : s.bipartite_above (swap r) b = s.bipartite_below r b := rfl +@[simp, norm_cast] lemma coe_bipartite_below : (s.bipartite_below r b : set α) = {a ∈ s | r a b} := +coe_filter _ _ + +@[simp, norm_cast] lemma coe_bipartite_above : (t.bipartite_above r a : set β) = {b ∈ t | r a b} := +coe_filter _ _ + variables {s t a a' b b'} @[simp] lemma mem_bipartite_below {a : α} : a ∈ s.bipartite_below r b ↔ a ∈ s ∧ r a b := mem_filter @@ -79,5 +87,33 @@ lemma card_mul_eq_card_mul [Π a b, decidable (r a b)] (card_mul_le_card_mul _ (λ a ha, (hm a ha).ge) $ λ b hb, (hn b hb).le).antisymm $ card_mul_le_card_mul' _ (λ a ha, (hn a ha).ge) $ λ b hb, (hm b hb).le +lemma card_le_card_of_forall_subsingleton + (hs : ∀ a ∈ s, ∃ b, b ∈ t ∧ r a b) (ht : ∀ b ∈ t, ({a ∈ s | r a b} : set α).subsingleton) : + s.card ≤ t.card := +by classical; simpa using card_mul_le_card_mul _ (λ a h, card_pos.2 $ + (by { rw [←coe_nonempty, coe_bipartite_above], exact hs _ h } : (t.bipartite_above r a).nonempty)) + (λ b h, card_le_one.2 $ by { simp_rw mem_bipartite_below, exact ht _ h }) + +lemma card_le_card_of_forall_subsingleton' + (ht : ∀ b ∈ t, ∃ a, a ∈ s ∧ r a b) (hs : ∀ a ∈ s, ({b ∈ t | r a b} : set β).subsingleton) : + t.card ≤ s.card := +card_le_card_of_forall_subsingleton (swap r) ht hs + end bipartite end finset + +open finset + +namespace fintype +variables [fintype α] [fintype β] {r : α → β → Prop} + +lemma card_le_card_of_left_total_unique (h₁ : left_total r) (h₂ : left_unique r) : + fintype.card α ≤ fintype.card β := +card_le_card_of_forall_subsingleton r (by simpa using h₁) $ λ b _ a₁ ha₁ a₂ ha₂, h₂ ha₁.2 ha₂.2 + +lemma card_le_card_of_right_total_unique (h₁ : right_total r) (h₂ : right_unique r) : + fintype.card β ≤ fintype.card α := +card_le_card_of_forall_subsingleton' r (by simpa using h₁) $ λ b _ a₁ ha₁ a₂ ha₂, h₂ ha₁.2 ha₂.2 + +end fintype + From daf8adf14245800e936d0b7f3c9890d259eac81e Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Wed, 14 Dec 2022 21:31:23 +0000 Subject: [PATCH 0048/1291] feat(algebra/star/basic): add ring_hom.has_involutive_star (#17904) Define an instance `has_involutive_star` for ring homorphisms into a `star_ring`. --- src/algebra/star/basic.lean | 12 ++++++++++++ src/algebra/star/self_adjoint.lean | 4 ++++ 2 files changed, 16 insertions(+) diff --git a/src/algebra/star/basic.lean b/src/algebra/star/basic.lean index 9e4be9cab8f6b..2229212f9d368 100644 --- a/src/algebra/star/basic.lean +++ b/src/algebra/star/basic.lean @@ -291,6 +291,18 @@ lemma star_ring_end_apply [comm_semiring R] [star_ring R] {x : R} : @[simp] lemma star_ring_end_self_apply [comm_semiring R] [star_ring R] (x : R) : star_ring_end R (star_ring_end R x) = x := star_star x +instance ring_hom.has_involutive_star {S : Type u} [non_assoc_semiring S] [comm_semiring R] + [star_ring R] : has_involutive_star (S →+* R) := +{ to_has_star := { star := λ f, ring_hom.comp (star_ring_end R) f }, + star_involutive := + by { intro _, ext, simp only [ring_hom.coe_comp, function.comp_app, star_ring_end_self_apply] }} + +lemma ring_hom.star_def {S : Type u} [non_assoc_semiring S] [comm_semiring R] [star_ring R] + (f : S →+* R) : has_star.star f = ring_hom.comp (star_ring_end R) f := rfl + +lemma ring_hom.star_apply {S : Type u} [non_assoc_semiring S] [comm_semiring R] [star_ring R] + (f : S →+* R) (s : S) : star f s = star (f s) := rfl + -- A more convenient name for complex conjugation alias star_ring_end_self_apply ← complex.conj_conj alias star_ring_end_self_apply ← is_R_or_C.conj_conj diff --git a/src/algebra/star/self_adjoint.lean b/src/algebra/star/self_adjoint.lean index d33bddca15b18..81aa4b3834f10 100644 --- a/src/algebra/star/self_adjoint.lean +++ b/src/algebra/star/self_adjoint.lean @@ -58,6 +58,10 @@ lemma star_eq [has_star R] {x : R} (hx : is_self_adjoint x) : star x = x := hx lemma _root_.is_self_adjoint_iff [has_star R] {x : R} : is_self_adjoint x ↔ star x = x := iff.rfl +@[simp] +lemma star_iff [has_involutive_star R] {x : R} : is_self_adjoint (star x) ↔ is_self_adjoint x := +by simpa only [is_self_adjoint, star_star] using eq_comm + @[simp] lemma star_mul_self [semigroup R] [star_semigroup R] (x : R) : is_self_adjoint (star x * x) := by simp only [is_self_adjoint, star_mul, star_star] From 2f7913bd4c5549ea4a8ec3db3b8d9a114b436af6 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 15 Dec 2022 00:09:40 +0000 Subject: [PATCH 0049/1291] fix(ci/add_ported_warnings): use a version of the action that does not error (#17936) See https://github.com/jitterbit/get-changed-files/issues/19#issuecomment-854657729 for details. --- .github/workflows/add_ported_warnings.yml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/.github/workflows/add_ported_warnings.yml b/.github/workflows/add_ported_warnings.yml index c620e17518c81..a01bef3721670 100644 --- a/.github/workflows/add_ported_warnings.yml +++ b/.github/workflows/add_ported_warnings.yml @@ -22,7 +22,7 @@ jobs: # TODO: is this really faster than just calling git from python? - name: Get changed files id: changed-files - uses: jitterbit/get-changed-files@v1 + uses: Ana06/get-changed-files@v1.2 - name: run the script run: | From 0efabe73557d6f5070d7ec68cab671a8a786ae7b Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Thu, 15 Dec 2022 01:17:53 +0000 Subject: [PATCH 0050/1291] fix(geometry/euclidean/angle/unoriented/right_angle): fix lemma naming (#17941) As pointed out by @eric-wieser in review of #17683 for corresponding oriented angle lemmas, some lemma names refer to `norm` but should refer to `dist`. --- .../euclidean/angle/unoriented/right_angle.lean | 12 ++++++------ 1 file changed, 6 insertions(+), 6 deletions(-) diff --git a/src/geometry/euclidean/angle/unoriented/right_angle.lean b/src/geometry/euclidean/angle/unoriented/right_angle.lean index 21e7df7841621..8c6a7b71c02e8 100644 --- a/src/geometry/euclidean/angle/unoriented/right_angle.lean +++ b/src/geometry/euclidean/angle/unoriented/right_angle.lean @@ -497,7 +497,7 @@ end /-- The cosine of an angle in a right-angled triangle multiplied by the hypotenuse equals the adjacent side. -/ -lemma cos_angle_mul_norm_of_angle_eq_pi_div_two {p₁ p₂ p₃ : P} (h : ∠ p₁ p₂ p₃ = π / 2) : +lemma cos_angle_mul_dist_of_angle_eq_pi_div_two {p₁ p₂ p₃ : P} (h : ∠ p₁ p₂ p₃ = π / 2) : real.cos (∠ p₂ p₃ p₁) * dist p₁ p₃ = dist p₃ p₂ := begin rw [angle, ←inner_eq_zero_iff_angle_eq_pi_div_two, real_inner_comm, ←neg_eq_zero, @@ -508,7 +508,7 @@ end /-- The sine of an angle in a right-angled triangle multiplied by the hypotenuse equals the opposite side. -/ -lemma sin_angle_mul_norm_of_angle_eq_pi_div_two {p₁ p₂ p₃ : P} (h : ∠ p₁ p₂ p₃ = π / 2) : +lemma sin_angle_mul_dist_of_angle_eq_pi_div_two {p₁ p₂ p₃ : P} (h : ∠ p₁ p₂ p₃ = π / 2) : real.sin (∠ p₂ p₃ p₁) * dist p₁ p₃ = dist p₁ p₂ := begin rw [angle, ←inner_eq_zero_iff_angle_eq_pi_div_two, real_inner_comm, ←neg_eq_zero, @@ -519,7 +519,7 @@ end /-- The tangent of an angle in a right-angled triangle multiplied by the adjacent side equals the opposite side. -/ -lemma tan_angle_mul_norm_of_angle_eq_pi_div_two {p₁ p₂ p₃ : P} (h : ∠ p₁ p₂ p₃ = π / 2) +lemma tan_angle_mul_dist_of_angle_eq_pi_div_two {p₁ p₂ p₃ : P} (h : ∠ p₁ p₂ p₃ = π / 2) (h0 : p₁ = p₂ ∨ p₃ ≠ p₂) : real.tan (∠ p₂ p₃ p₁) * dist p₃ p₂ = dist p₁ p₂ := begin rw [angle, ←inner_eq_zero_iff_angle_eq_pi_div_two, real_inner_comm, ←neg_eq_zero, @@ -531,7 +531,7 @@ end /-- A side of a right-angled triangle divided by the cosine of the adjacent angle equals the hypotenuse. -/ -lemma norm_div_cos_angle_of_angle_eq_pi_div_two {p₁ p₂ p₃ : P} (h : ∠ p₁ p₂ p₃ = π / 2) +lemma dist_div_cos_angle_of_angle_eq_pi_div_two {p₁ p₂ p₃ : P} (h : ∠ p₁ p₂ p₃ = π / 2) (h0 : p₁ = p₂ ∨ p₃ ≠ p₂) : dist p₃ p₂ / real.cos (∠ p₂ p₃ p₁) = dist p₁ p₃ := begin rw [angle, ←inner_eq_zero_iff_angle_eq_pi_div_two, real_inner_comm, ←neg_eq_zero, @@ -543,7 +543,7 @@ end /-- A side of a right-angled triangle divided by the sine of the opposite angle equals the hypotenuse. -/ -lemma norm_div_sin_angle_of_angle_eq_pi_div_two {p₁ p₂ p₃ : P} (h : ∠ p₁ p₂ p₃ = π / 2) +lemma dist_div_sin_angle_of_angle_eq_pi_div_two {p₁ p₂ p₃ : P} (h : ∠ p₁ p₂ p₃ = π / 2) (h0 : p₁ ≠ p₂ ∨ p₃ = p₂) : dist p₁ p₂ / real.sin (∠ p₂ p₃ p₁) = dist p₁ p₃ := begin rw [angle, ←inner_eq_zero_iff_angle_eq_pi_div_two, real_inner_comm, ←neg_eq_zero, @@ -555,7 +555,7 @@ end /-- A side of a right-angled triangle divided by the tangent of the opposite angle equals the adjacent side. -/ -lemma norm_div_tan_angle_of_angle_eq_pi_div_two {p₁ p₂ p₃ : P} (h : ∠ p₁ p₂ p₃ = π / 2) +lemma dist_div_tan_angle_of_angle_eq_pi_div_two {p₁ p₂ p₃ : P} (h : ∠ p₁ p₂ p₃ = π / 2) (h0 : p₁ ≠ p₂ ∨ p₃ = p₂) : dist p₁ p₂ / real.tan (∠ p₂ p₃ p₁) = dist p₃ p₂ := begin rw [angle, ←inner_eq_zero_iff_angle_eq_pi_div_two, real_inner_comm, ←neg_eq_zero, From a59dad53320b73ef180174aae867addd707ef00e Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Thu, 15 Dec 2022 04:23:57 +0000 Subject: [PATCH 0051/1291] chore(*): add mathlib4 synchronization comments (#17943) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following PRs: * `algebra.euclidean_domain.basic`: https://github.com/leanprover-community/mathlib4/pull/919 * `algebra.field.basic`: https://github.com/leanprover-community/mathlib4/pull/975 * `algebra.group.opposite`: https://github.com/leanprover-community/mathlib4/pull/912 * `algebra.group.prod`: https://github.com/leanprover-community/mathlib4/pull/968 * `algebra.group.with_one.units`: https://github.com/leanprover-community/mathlib4/pull/955 * `algebra.group_power.ring`: https://github.com/leanprover-community/mathlib4/pull/979 * `algebra.hom.equiv.type_tags`: https://github.com/leanprover-community/mathlib4/pull/943 * `algebra.hom.ring`: https://github.com/leanprover-community/mathlib4/pull/958 * `algebra.invertible`: https://github.com/leanprover-community/mathlib4/pull/930 * `algebra.order.field.canonical.basic`: https://github.com/leanprover-community/mathlib4/pull/1018 * `algebra.order.field.canonical.defs`: https://github.com/leanprover-community/mathlib4/pull/985 * `algebra.order.field.inj_surj`: https://github.com/leanprover-community/mathlib4/pull/1017 * `algebra.order.group.densely_ordered`: https://github.com/leanprover-community/mathlib4/pull/956 * `algebra.order.group.min_max`: https://github.com/leanprover-community/mathlib4/pull/933 * `algebra.order.group.prod`: https://github.com/leanprover-community/mathlib4/pull/1026 * `algebra.order.group.with_top`: https://github.com/leanprover-community/mathlib4/pull/946 * `algebra.order.hom.monoid`: https://github.com/leanprover-community/mathlib4/pull/944 * `algebra.order.monoid.prod`: https://github.com/leanprover-community/mathlib4/pull/987 * `algebra.order.monoid.to_mul_bot`: https://github.com/leanprover-community/mathlib4/pull/1024 * `algebra.order.ring.abs`: https://github.com/leanprover-community/mathlib4/pull/929 * `algebra.order.ring.cone`: https://github.com/leanprover-community/mathlib4/pull/935 * `algebra.order.sub.basic`: https://github.com/leanprover-community/mathlib4/pull/936 * `algebra.order.sub.with_top`: https://github.com/leanprover-community/mathlib4/pull/932 * `algebra.order.with_zero`: https://github.com/leanprover-community/mathlib4/pull/903 * `combinatorics.quiver.arborescence`: https://github.com/leanprover-community/mathlib4/pull/997 * `combinatorics.quiver.cast`: https://github.com/leanprover-community/mathlib4/pull/990 * `control.traversable.lemmas`: https://github.com/leanprover-community/mathlib4/pull/948 * `data.bool.set`: https://github.com/leanprover-community/mathlib4/pull/960 * `data.int.cast.field`: https://github.com/leanprover-community/mathlib4/pull/1016 * `data.int.cast.lemmas`: https://github.com/leanprover-community/mathlib4/pull/995 * `data.int.cast.prod`: https://github.com/leanprover-community/mathlib4/pull/1015 * `data.int.div`: https://github.com/leanprover-community/mathlib4/pull/1011 * `data.int.dvd.basic`: https://github.com/leanprover-community/mathlib4/pull/996 * `data.int.order.basic`: https://github.com/leanprover-community/mathlib4/pull/938 * `data.nat.cast.basic`: https://github.com/leanprover-community/mathlib4/pull/980 * `data.nat.cast.prod`: https://github.com/leanprover-community/mathlib4/pull/1010 * `data.nat.cast.with_top`: https://github.com/leanprover-community/mathlib4/pull/1019 * `data.nat.gcd.basic`: https://github.com/leanprover-community/mathlib4/pull/965 * `data.nat.order.basic`: https://github.com/leanprover-community/mathlib4/pull/907 * `data.nat.order.lemmas`: https://github.com/leanprover-community/mathlib4/pull/927 * `data.nat.set`: https://github.com/leanprover-community/mathlib4/pull/961 * `data.nat.upto`: https://github.com/leanprover-community/mathlib4/pull/1020 * `data.pequiv`: https://github.com/leanprover-community/mathlib4/pull/1008 * `data.set.basic`: https://github.com/leanprover-community/mathlib4/pull/892 * `data.set.bool_indicator`: https://github.com/leanprover-community/mathlib4/pull/988 * `data.set.image`: https://github.com/leanprover-community/mathlib4/pull/949 * `data.set.n_ary`: https://github.com/leanprover-community/mathlib4/pull/969 * `data.set.opposite`: https://github.com/leanprover-community/mathlib4/pull/983 * `data.set.prod`: https://github.com/leanprover-community/mathlib4/pull/1004 * `data.set.sigma`: https://github.com/leanprover-community/mathlib4/pull/982 * `data.set_like.basic`: https://github.com/leanprover-community/mathlib4/pull/993 * `data.two_pointing`: https://github.com/leanprover-community/mathlib4/pull/984 * `logic.embedding.set`: https://github.com/leanprover-community/mathlib4/pull/986 * `logic.equiv.embedding`: https://github.com/leanprover-community/mathlib4/pull/1021 * `order.directed`: https://github.com/leanprover-community/mathlib4/pull/963 * `order.rel_iso.set`: https://github.com/leanprover-community/mathlib4/pull/1005 * `order.well_founded`: https://github.com/leanprover-community/mathlib4/pull/970 * `ring_theory.coprime.basic`: https://github.com/leanprover-community/mathlib4/pull/899 Co-authored-by: Jireh Loreaux --- src/algebra/euclidean_domain/basic.lean | 4 ++++ src/algebra/field/basic.lean | 4 ++++ src/algebra/group/opposite.lean | 4 ++++ src/algebra/group/prod.lean | 4 ++++ src/algebra/group/with_one/units.lean | 4 ++++ src/algebra/group_power/ring.lean | 4 ++++ src/algebra/hom/equiv/type_tags.lean | 4 ++++ src/algebra/hom/ring.lean | 4 ++++ src/algebra/invertible.lean | 4 ++++ src/algebra/order/field/canonical/basic.lean | 4 ++++ src/algebra/order/field/canonical/defs.lean | 4 ++++ src/algebra/order/field/inj_surj.lean | 4 ++++ src/algebra/order/group/densely_ordered.lean | 4 ++++ src/algebra/order/group/min_max.lean | 4 ++++ src/algebra/order/group/prod.lean | 4 ++++ src/algebra/order/group/with_top.lean | 4 ++++ src/algebra/order/hom/monoid.lean | 4 ++++ src/algebra/order/monoid/prod.lean | 6 +++++- src/algebra/order/monoid/to_mul_bot.lean | 4 ++++ src/algebra/order/ring/abs.lean | 4 ++++ src/algebra/order/ring/cone.lean | 4 ++++ src/algebra/order/sub/basic.lean | 4 ++++ src/algebra/order/sub/with_top.lean | 4 ++++ src/algebra/order/with_zero.lean | 4 ++++ src/combinatorics/quiver/arborescence.lean | 4 ++++ src/combinatorics/quiver/cast.lean | 4 ++++ src/control/traversable/lemmas.lean | 4 ++++ src/data/bool/set.lean | 4 ++++ src/data/int/cast/field.lean | 4 ++++ src/data/int/cast/lemmas.lean | 4 ++++ src/data/int/cast/prod.lean | 4 ++++ src/data/int/div.lean | 4 ++++ src/data/int/dvd/basic.lean | 4 ++++ src/data/int/order/basic.lean | 4 ++++ src/data/nat/cast/basic.lean | 4 ++++ src/data/nat/cast/prod.lean | 4 ++++ src/data/nat/cast/with_top.lean | 4 ++++ src/data/nat/gcd/basic.lean | 4 ++++ src/data/nat/order/basic.lean | 4 ++++ src/data/nat/order/lemmas.lean | 4 ++++ src/data/nat/set.lean | 4 ++++ src/data/nat/upto.lean | 4 ++++ src/data/pequiv.lean | 4 ++++ src/data/set/basic.lean | 4 ++++ src/data/set/bool_indicator.lean | 4 ++++ src/data/set/image.lean | 4 ++++ src/data/set/n_ary.lean | 4 ++++ src/data/set/opposite.lean | 4 ++++ src/data/set/prod.lean | 4 ++++ src/data/set/sigma.lean | 4 ++++ src/data/set_like/basic.lean | 4 ++++ src/data/two_pointing.lean | 4 ++++ src/logic/embedding/set.lean | 4 ++++ src/logic/equiv/embedding.lean | 4 ++++ src/order/directed.lean | 4 ++++ src/order/rel_iso/set.lean | 4 ++++ src/order/well_founded.lean | 4 ++++ src/ring_theory/coprime/basic.lean | 4 ++++ 58 files changed, 233 insertions(+), 1 deletion(-) diff --git a/src/algebra/euclidean_domain/basic.lean b/src/algebra/euclidean_domain/basic.lean index 6a9c99df93b0f..045ef77015bce 100644 --- a/src/algebra/euclidean_domain/basic.lean +++ b/src/algebra/euclidean_domain/basic.lean @@ -12,6 +12,10 @@ import algebra.ring.basic /-! # Lemmas about Euclidean domains +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/919 +> Any changes to this file require a corresponding PR to mathlib4. + ## Main statements * `gcd_eq_gcd_ab`: states Bézout's lemma for Euclidean domains. diff --git a/src/algebra/field/basic.lean b/src/algebra/field/basic.lean index 4dde0718e8d15..43b82b961b4c0 100644 --- a/src/algebra/field/basic.lean +++ b/src/algebra/field/basic.lean @@ -11,6 +11,10 @@ import algebra.ring.inj_surj /-! # Lemmas about division (semi)rings and (semi)fields +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/975 +> Any changes to this file require a corresponding PR to mathlib4. + -/ open function order_dual set diff --git a/src/algebra/group/opposite.lean b/src/algebra/group/opposite.lean index ee00fdab71b3d..75b4412a2e0c2 100644 --- a/src/algebra/group/opposite.lean +++ b/src/algebra/group/opposite.lean @@ -11,6 +11,10 @@ import data.int.cast.defs /-! # Group structures on the multiplicative and additive opposites + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/912 +> Any changes to this file require a corresponding PR to mathlib4. -/ universes u v variables (α : Type u) diff --git a/src/algebra/group/prod.lean b/src/algebra/group/prod.lean index 0e68886b5f7e1..dbbb9c011d845 100644 --- a/src/algebra/group/prod.lean +++ b/src/algebra/group/prod.lean @@ -10,6 +10,10 @@ import algebra.hom.units /-! # Monoid, group etc structures on `M × N` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/968 +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define one-binop (`monoid`, `group` etc) structures on `M × N`. We also prove trivial `simp` lemmas, and define the following operations on `monoid_hom`s: diff --git a/src/algebra/group/with_one/units.lean b/src/algebra/group/with_one/units.lean index 41aff905b5f1f..76a2b5e139891 100644 --- a/src/algebra/group/with_one/units.lean +++ b/src/algebra/group/with_one/units.lean @@ -9,6 +9,10 @@ import algebra.group_with_zero.units.basic /-! # Isomorphism between a group and the units of itself adjoined with `0` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/955 +> Any changes to this file require a corresponding PR to mathlib4. + ## Notes This is here to keep `algebra.group_with_zero.units.basic` out of the import requirements of `algebra.order.field.defs`. diff --git a/src/algebra/group_power/ring.lean b/src/algebra/group_power/ring.lean index 5ddc0a56919e7..a774790640173 100644 --- a/src/algebra/group_power/ring.lean +++ b/src/algebra/group_power/ring.lean @@ -14,6 +14,10 @@ import data.nat.order.basic /-! # Power operations on monoids with zero, semirings, and rings +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/979 +> Any changes to this file require a corresponding PR to mathlib4. + This file provides additional lemmas about the natural power operator on rings and semirings. Further lemmas about ordered semirings and rings can be found in `algebra.group_power.lemmas`. diff --git a/src/algebra/hom/equiv/type_tags.lean b/src/algebra/hom/equiv/type_tags.lean index 9f40a2cd08224..329c5a2543094 100644 --- a/src/algebra/hom/equiv/type_tags.lean +++ b/src/algebra/hom/equiv/type_tags.lean @@ -8,6 +8,10 @@ import algebra.group.type_tags /-! # Additive and multiplicative equivalences associated to `multiplicative` and `additive`. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/943 +> Any changes to this file require a corresponding PR to mathlib4. -/ variables {G H : Type*} diff --git a/src/algebra/hom/ring.lean b/src/algebra/hom/ring.lean index ebf2272eeeab3..f2d7b2be544ff 100644 --- a/src/algebra/hom/ring.lean +++ b/src/algebra/hom/ring.lean @@ -13,6 +13,10 @@ import data.set.image /-! # Homomorphisms of semirings and rings +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/958 +> Any changes to this file require a corresponding PR to mathlib4. + This file defines bundled homomorphisms of (non-unital) semirings and rings. As with monoid and groups, we use the same structure `ring_hom a β`, a.k.a. `α →+* β`, for both types of homomorphisms. diff --git a/src/algebra/invertible.lean b/src/algebra/invertible.lean index 3c9005ec6209b..606c3401e78e5 100644 --- a/src/algebra/invertible.lean +++ b/src/algebra/invertible.lean @@ -11,6 +11,10 @@ import algebra.ring.defs /-! # Invertible elements +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/930 +> Any changes to this file require a corresponding PR to mathlib4. + This file defines a typeclass `invertible a` for elements `a` with a two-sided multiplicative inverse. diff --git a/src/algebra/order/field/canonical/basic.lean b/src/algebra/order/field/canonical/basic.lean index d37ac79c840bc..57dd84c1459f7 100644 --- a/src/algebra/order/field/canonical/basic.lean +++ b/src/algebra/order/field/canonical/basic.lean @@ -8,6 +8,10 @@ import algebra.order.field.canonical.defs /-! # Lemmas about canonically ordered semifields. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/1018 +> Any changes to this file require a corresponding PR to mathlib4. + -/ variables {α : Type*} diff --git a/src/algebra/order/field/canonical/defs.lean b/src/algebra/order/field/canonical/defs.lean index 6cfa773e4c813..674b364673911 100644 --- a/src/algebra/order/field/canonical/defs.lean +++ b/src/algebra/order/field/canonical/defs.lean @@ -10,6 +10,10 @@ import algebra.order.with_zero /-! # Canonically ordered semifields +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/985 +> Any changes to this file require a corresponding PR to mathlib4. + -/ set_option old_structure_cmd true diff --git a/src/algebra/order/field/inj_surj.lean b/src/algebra/order/field/inj_surj.lean index dcdde29d91008..a9c34901d1726 100644 --- a/src/algebra/order/field/inj_surj.lean +++ b/src/algebra/order/field/inj_surj.lean @@ -10,6 +10,10 @@ import algebra.order.ring.inj_surj /-! # Pulling back linearly ordered fields along injective maps. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/1017 +> Any changes to this file require a corresponding PR to mathlib4. + -/ open function order_dual diff --git a/src/algebra/order/group/densely_ordered.lean b/src/algebra/order/group/densely_ordered.lean index 822613485feba..2d2bb5f7cd002 100644 --- a/src/algebra/order/group/densely_ordered.lean +++ b/src/algebra/order/group/densely_ordered.lean @@ -9,6 +9,10 @@ import algebra.order.monoid.order_dual /-! # Lemmas about densely linearly ordered groups. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/956 +> Any changes to this file require a corresponding PR to mathlib4. -/ variables {α : Type*} diff --git a/src/algebra/order/group/min_max.lean b/src/algebra/order/group/min_max.lean index adefb7abbb2ca..9f0cd8d5c1d6a 100644 --- a/src/algebra/order/group/min_max.lean +++ b/src/algebra/order/group/min_max.lean @@ -8,6 +8,10 @@ import algebra.order.monoid.min_max /-! # `min` and `max` in linearly ordered groups. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/933 +> Any changes to this file require a corresponding PR to mathlib4. -/ section diff --git a/src/algebra/order/group/prod.lean b/src/algebra/order/group/prod.lean index 4a255df846e42..1b0639868d086 100644 --- a/src/algebra/order/group/prod.lean +++ b/src/algebra/order/group/prod.lean @@ -8,6 +8,10 @@ import algebra.order.monoid.prod /-! # Products of ordered commutative groups. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/1026 +> Any changes to this file require a corresponding PR to mathlib4. -/ variable {α : Type*} diff --git a/src/algebra/order/group/with_top.lean b/src/algebra/order/group/with_top.lean index f0fc8553bcfbf..6f23fc042b36b 100644 --- a/src/algebra/order/group/with_top.lean +++ b/src/algebra/order/group/with_top.lean @@ -8,6 +8,10 @@ import algebra.order.monoid.with_top /-! # Adjoining a top element to a `linear_ordered_add_comm_group_with_top`. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/946 +> Any changes to this file require a corresponding PR to mathlib4. -/ variable {α : Type*} diff --git a/src/algebra/order/hom/monoid.lean b/src/algebra/order/hom/monoid.lean index 10859efb337d8..c3e866f6e61b0 100644 --- a/src/algebra/order/hom/monoid.lean +++ b/src/algebra/order/hom/monoid.lean @@ -12,6 +12,10 @@ import order.hom.basic /-! # Ordered monoid and group homomorphisms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/944 +> Any changes to this file require a corresponding PR to mathlib4. + This file defines morphisms between (additive) ordered monoids. ## Types of morphisms diff --git a/src/algebra/order/monoid/prod.lean b/src/algebra/order/monoid/prod.lean index e7f0d95f7729e..e632cf805920a 100644 --- a/src/algebra/order/monoid/prod.lean +++ b/src/algebra/order/monoid/prod.lean @@ -7,7 +7,11 @@ import algebra.group.prod import algebra.order.monoid.cancel.defs import algebra.order.monoid.canonical.defs -/-! # Products of ordered monoids -/ +/-! # Products of ordered monoids + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/987 +> Any changes to this file require a corresponding PR to mathlib4.-/ namespace prod diff --git a/src/algebra/order/monoid/to_mul_bot.lean b/src/algebra/order/monoid/to_mul_bot.lean index 87b6ec8497284..e1f8ea61f6e7c 100644 --- a/src/algebra/order/monoid/to_mul_bot.lean +++ b/src/algebra/order/monoid/to_mul_bot.lean @@ -8,6 +8,10 @@ import algebra.order.monoid.with_top import algebra.order.monoid.type_tags /-! +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/1024 +> Any changes to this file require a corresponding PR to mathlib4. + Making an additive monoid multiplicative then adding a zero is the same as adding a bottom element then making it multiplicative. -/ diff --git a/src/algebra/order/ring/abs.lean b/src/algebra/order/ring/abs.lean index 05e6c1a7efc11..24f8dca1b4f8c 100644 --- a/src/algebra/order/ring/abs.lean +++ b/src/algebra/order/ring/abs.lean @@ -10,6 +10,10 @@ import algebra.order.group.abs /-! # Absolute values in linear ordered rings. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/929 +> Any changes to this file require a corresponding PR to mathlib4. -/ variables {α : Type*} diff --git a/src/algebra/order/ring/cone.lean b/src/algebra/order/ring/cone.lean index a07fdc0f99c79..212f254206158 100644 --- a/src/algebra/order/ring/cone.lean +++ b/src/algebra/order/ring/cone.lean @@ -8,6 +8,10 @@ import algebra.order.ring.defs /-! # Constructing an ordered ring from a ring with a specified positive cone. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/935 +> Any changes to this file require a corresponding PR to mathlib4. + -/ /-! ### Positive cones -/ diff --git a/src/algebra/order/sub/basic.lean b/src/algebra/order/sub/basic.lean index 35a78f620700f..f4ca27340619c 100644 --- a/src/algebra/order/sub/basic.lean +++ b/src/algebra/order/sub/basic.lean @@ -11,6 +11,10 @@ import algebra.order.sub.defs /-! # Additional results about ordered Subtraction +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/936 +> Any changes to this file require a corresponding PR to mathlib4. + -/ variables {α β : Type*} diff --git a/src/algebra/order/sub/with_top.lean b/src/algebra/order/sub/with_top.lean index 9068932b60cd7..82589c1cfb386 100644 --- a/src/algebra/order/sub/with_top.lean +++ b/src/algebra/order/sub/with_top.lean @@ -8,6 +8,10 @@ import algebra.order.monoid.with_top /-! # Lemma about subtraction in ordered monoids with a top element adjoined. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/932 +> Any changes to this file require a corresponding PR to mathlib4. -/ variables {α β : Type*} diff --git a/src/algebra/order/with_zero.lean b/src/algebra/order/with_zero.lean index 8d5df167e102d..02133e056250f 100644 --- a/src/algebra/order/with_zero.lean +++ b/src/algebra/order/with_zero.lean @@ -14,6 +14,10 @@ import algebra.order.monoid.type_tags /-! # Linearly ordered commutative groups and monoids with a zero element adjoined +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/903 +> Any changes to this file require a corresponding PR to mathlib4. + This file sets up a special class of linearly ordered commutative monoids that show up as the target of so-called “valuations” in algebraic number theory. diff --git a/src/combinatorics/quiver/arborescence.lean b/src/combinatorics/quiver/arborescence.lean index 72b384ba9854b..923b7fd149f16 100644 --- a/src/combinatorics/quiver/arborescence.lean +++ b/src/combinatorics/quiver/arborescence.lean @@ -11,6 +11,10 @@ import combinatorics.quiver.path /-! # Arborescences +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/997 +> Any changes to this file require a corresponding PR to mathlib4. + A quiver `V` is an arborescence (or directed rooted tree) when we have a root vertex `root : V` such that for every `b : V` there is a unique path from `root` to `b`. diff --git a/src/combinatorics/quiver/cast.lean b/src/combinatorics/quiver/cast.lean index 3687801c5617e..d4a2de4e3bae9 100644 --- a/src/combinatorics/quiver/cast.lean +++ b/src/combinatorics/quiver/cast.lean @@ -10,6 +10,10 @@ import combinatorics.quiver.path # Rewriting arrows and paths along vertex equalities +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/990 +> Any changes to this file require a corresponding PR to mathlib4. + This files defines `hom.cast` and `path.cast` (and associated lemmas) in order to allow rewriting arrows and paths along equalities of their endpoints. diff --git a/src/control/traversable/lemmas.lean b/src/control/traversable/lemmas.lean index 2a4134ddcbfdc..87e4df83e52d4 100644 --- a/src/control/traversable/lemmas.lean +++ b/src/control/traversable/lemmas.lean @@ -9,6 +9,10 @@ import control.traversable.basic /-! # Traversing collections +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/948 +> Any changes to this file require a corresponding PR to mathlib4. + This file proves basic properties of traversable and applicative functors and defines `pure_transformation F`, the natural applicative transformation from the identity functor to `F`. diff --git a/src/data/bool/set.lean b/src/data/bool/set.lean index bf4558a70864b..8128f19f96d33 100644 --- a/src/data/bool/set.lean +++ b/src/data/bool/set.lean @@ -9,6 +9,10 @@ import data.set.image /-! # Booleans and set operations +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/960 +> Any changes to this file require a corresponding PR to mathlib4. + This file contains two trivial lemmas about `bool`, `set.univ`, and `set.range`. -/ diff --git a/src/data/int/cast/field.lean b/src/data/int/cast/field.lean index b96ce2bc5415c..ac4d838b73687 100644 --- a/src/data/int/cast/field.lean +++ b/src/data/int/cast/field.lean @@ -10,6 +10,10 @@ import algebra.group_with_zero.units.lemmas /-! # Cast of integers into fields +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/1016 +> Any changes to this file require a corresponding PR to mathlib4. + This file concerns the canonical homomorphism `ℤ → F`, where `F` is a field. ## Main results diff --git a/src/data/int/cast/lemmas.lean b/src/data/int/cast/lemmas.lean index bd6c39f68e890..a08b6ceed64ac 100644 --- a/src/data/int/cast/lemmas.lean +++ b/src/data/int/cast/lemmas.lean @@ -9,6 +9,10 @@ import data.nat.cast.basic /-! # Cast of integers (additional theorems) +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/995 +> Any changes to this file require a corresponding PR to mathlib4. + This file proves additional properties about the *canonical* homomorphism from the integers into an additive group with a one (`int.cast`), particularly results involving algebraic homomorphisms or the order structure on `ℤ` diff --git a/src/data/int/cast/prod.lean b/src/data/int/cast/prod.lean index 8482a3857906b..c4b1897f91a61 100644 --- a/src/data/int/cast/prod.lean +++ b/src/data/int/cast/prod.lean @@ -8,6 +8,10 @@ import data.nat.cast.prod /-! # The product of two `add_group_with_one`s. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/1015 +> Any changes to this file require a corresponding PR to mathlib4. -/ namespace prod diff --git a/src/data/int/div.lean b/src/data/int/div.lean index 3a93b34352da6..4d0724f1449bb 100644 --- a/src/data/int/div.lean +++ b/src/data/int/div.lean @@ -9,6 +9,10 @@ import algebra.ring.regular /-! # Lemmas relating `/` in `ℤ` with the ordering. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/1011 +> Any changes to this file require a corresponding PR to mathlib4. -/ open nat diff --git a/src/data/int/dvd/basic.lean b/src/data/int/dvd/basic.lean index 2905eb1a1494e..24c7c5c92ced2 100644 --- a/src/data/int/dvd/basic.lean +++ b/src/data/int/dvd/basic.lean @@ -8,6 +8,10 @@ import data.nat.cast.basic /-! # Basic lemmas about the divisibility relation in `ℤ`. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/996 +> Any changes to this file require a corresponding PR to mathlib4. -/ open nat diff --git a/src/data/int/order/basic.lean b/src/data/int/order/basic.lean index 2bbabc6993ae0..d91f7a5b7f6d1 100644 --- a/src/data/int/order/basic.lean +++ b/src/data/int/order/basic.lean @@ -12,6 +12,10 @@ import tactic.assert_exists /-! # Order instances on the integers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/938 +> Any changes to this file require a corresponding PR to mathlib4. + This file contains: * instances on `ℤ`. The stronger one is `int.linear_ordered_comm_ring`. * basic lemmas about integers that involve order properties. diff --git a/src/data/nat/cast/basic.lean b/src/data/nat/cast/basic.lean index 3025b8b3e6d8e..56c3f44f5730a 100644 --- a/src/data/nat/cast/basic.lean +++ b/src/data/nat/cast/basic.lean @@ -14,6 +14,10 @@ import algebra.group.opposite /-! # Cast of natural numbers (additional theorems) +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/980 +> Any changes to this file require a corresponding PR to mathlib4. + This file proves additional properties about the *canonical* homomorphism from the natural numbers into an additive monoid with a one (`nat.cast`). diff --git a/src/data/nat/cast/prod.lean b/src/data/nat/cast/prod.lean index 4c218abee4acd..fa8de3dd8438b 100644 --- a/src/data/nat/cast/prod.lean +++ b/src/data/nat/cast/prod.lean @@ -8,6 +8,10 @@ import algebra.group.prod /-! # The product of two `add_monoid_with_one`s. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/1010 +> Any changes to this file require a corresponding PR to mathlib4. -/ variables {α β : Type*} diff --git a/src/data/nat/cast/with_top.lean b/src/data/nat/cast/with_top.lean index d17168e115c7a..ddd74be02e52f 100644 --- a/src/data/nat/cast/with_top.lean +++ b/src/data/nat/cast/with_top.lean @@ -9,6 +9,10 @@ import data.nat.basic /-! # Lemma about the coercion `ℕ → with_bot ℕ`. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/1019 +> Any changes to this file require a corresponding PR to mathlib4. + An orphaned lemma about casting from `ℕ` to `with_bot ℕ`, exiled here to minimize imports to `data.rat.order` for porting purposes. -/ diff --git a/src/data/nat/gcd/basic.lean b/src/data/nat/gcd/basic.lean index 36a748926949e..de73719b823cd 100644 --- a/src/data/nat/gcd/basic.lean +++ b/src/data/nat/gcd/basic.lean @@ -10,6 +10,10 @@ import data.nat.order.lemmas /-! # Definitions and properties of `nat.gcd`, `nat.lcm`, and `nat.coprime` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/965 +> Any changes to this file require a corresponding PR to mathlib4. + Generalizations of these are provided in a later file as `gcd_monoid.gcd` and `gcd_monoid.lcm`. diff --git a/src/data/nat/order/basic.lean b/src/data/nat/order/basic.lean index dea794afb9224..39e36ad9fc3ad 100644 --- a/src/data/nat/order/basic.lean +++ b/src/data/nat/order/basic.lean @@ -9,6 +9,10 @@ import data.nat.basic /-! # The natural numbers as a linearly ordered commutative semiring +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/907 +> Any changes to this file require a corresponding PR to mathlib4. + We also have a variety of lemmas which have been deferred from `data.nat.basic` because it is easier to prove them with this ordered semiring instance available. diff --git a/src/data/nat/order/lemmas.lean b/src/data/nat/order/lemmas.lean index b9e86827bdd55..8f244741fd277 100644 --- a/src/data/nat/order/lemmas.lean +++ b/src/data/nat/order/lemmas.lean @@ -11,6 +11,10 @@ import algebra.group_with_zero.divisibility /-! # Further lemmas about the natural numbers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/927 +> Any changes to this file require a corresponding PR to mathlib4. + The distinction between this file and `data.nat.order.basic` is not particularly clear. They are separated by now to minimize the porting requirements for tactics during the transition to mathlib4. After `data.rat.order` has been ported, please feel free to reorganize these two files. diff --git a/src/data/nat/set.lean b/src/data/nat/set.lean index 6d372cb3be719..7e829522e8e34 100644 --- a/src/data/nat/set.lean +++ b/src/data/nat/set.lean @@ -8,6 +8,10 @@ import data.set.image /-! ### Recursion on the natural numbers and `set.range` + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/961 +> Any changes to this file require a corresponding PR to mathlib4. -/ namespace nat diff --git a/src/data/nat/upto.lean b/src/data/nat/upto.lean index e8996cd788fc8..5a4322ecf74dd 100644 --- a/src/data/nat/upto.lean +++ b/src/data/nat/upto.lean @@ -8,6 +8,10 @@ import data.nat.order.basic /-! # `nat.upto` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/1020 +> Any changes to this file require a corresponding PR to mathlib4. + `nat.upto p`, with `p` a predicate on `ℕ`, is a subtype of elements `n : ℕ` such that no value (strictly) below `n` satisfies `p`. diff --git a/src/data/pequiv.lean b/src/data/pequiv.lean index ea1cf4cd9b40d..e1451e5f8c07e 100644 --- a/src/data/pequiv.lean +++ b/src/data/pequiv.lean @@ -9,6 +9,10 @@ import data.set.basic # Partial Equivalences +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/1008 +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we define partial equivalences `pequiv`, which are a bijection between a subset of `α` and a subset of `β`. Notationally, a `pequiv` is denoted by "`≃.`" (note that the full stop is part of the notation). The way we store these internally is with two functions `f : α → option β` and diff --git a/src/data/set/basic.lean b/src/data/set/basic.lean index 24571fb1e9f51..97feb24603421 100644 --- a/src/data/set/basic.lean +++ b/src/data/set/basic.lean @@ -9,6 +9,10 @@ import logic.function.iterate /-! # Basic properties of sets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/892 +> Any changes to this file require a corresponding PR to mathlib4. + Sets in Lean are homogeneous; all their elements have the same type. Sets whose elements have type `X` are thus defined as `set X := X → Prop`. Note that this function need not be decidable. The definition is in the core library. diff --git a/src/data/set/bool_indicator.lean b/src/data/set/bool_indicator.lean index afbd45eadb363..3b95d84c1a935 100644 --- a/src/data/set/bool_indicator.lean +++ b/src/data/set/bool_indicator.lean @@ -9,6 +9,10 @@ import data.set.image /-! # Indicator function valued in bool +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/988 +> Any changes to this file require a corresponding PR to mathlib4. + See also `set.indicator` and `set.piecewise`. -/ diff --git a/src/data/set/image.lean b/src/data/set/image.lean index c89fcc1f6ec0a..7b49083f2bbf5 100644 --- a/src/data/set/image.lean +++ b/src/data/set/image.lean @@ -8,6 +8,10 @@ import data.set.basic /-! # Images and preimages of sets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/949 +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions * `preimage f t : set α` : the preimage f⁻¹(t) (written `f ⁻¹' t` in Lean) of a subset of β. diff --git a/src/data/set/n_ary.lean b/src/data/set/n_ary.lean index 5c26c45a1cefd..6eae6cc3b6262 100644 --- a/src/data/set/n_ary.lean +++ b/src/data/set/n_ary.lean @@ -8,6 +8,10 @@ import data.set.prod /-! # N-ary images of sets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/969 +> Any changes to this file require a corresponding PR to mathlib4. + 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. diff --git a/src/data/set/opposite.lean b/src/data/set/opposite.lean index a4f90f3ce7217..4f5f97cd4f488 100644 --- a/src/data/set/opposite.lean +++ b/src/data/set/opposite.lean @@ -9,6 +9,10 @@ import data.set.image /-! # The opposite of a set +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/983 +> Any changes to this file require a corresponding PR to mathlib4. + The opposite of a set `s` is simply the set obtained by taking the opposite of each member of `s`. -/ diff --git a/src/data/set/prod.lean b/src/data/set/prod.lean index 37c51a4e8ea71..c8f3bc183cefb 100644 --- a/src/data/set/prod.lean +++ b/src/data/set/prod.lean @@ -8,6 +8,10 @@ import data.set.image /-! # Sets in product and pi types +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/1004 +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the product of sets in `α × β` and in `Π i, α i` along with the diagonal of a type. diff --git a/src/data/set/sigma.lean b/src/data/set/sigma.lean index 1aa3bfb920dd8..24053c92aa2f2 100644 --- a/src/data/set/sigma.lean +++ b/src/data/set/sigma.lean @@ -8,6 +8,10 @@ import data.set.image /-! # Sets in sigma types +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/982 +> Any changes to this file require a corresponding PR to mathlib4. + This file defines `set.sigma`, the indexed sum of sets. -/ diff --git a/src/data/set_like/basic.lean b/src/data/set_like/basic.lean index 2176811dcb11f..c89d5dfc8733e 100644 --- a/src/data/set_like/basic.lean +++ b/src/data/set_like/basic.lean @@ -9,6 +9,10 @@ import tactic.monotonicity.basic /-! # Typeclass for types with a set-like extensionality property +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/993 +> Any changes to this file require a corresponding PR to mathlib4. + The `has_mem` typeclass is used to let terms of a type have elements. Many instances of `has_mem` have a set-like extensionality property: things are equal iff they have the same elements. The `set_like` diff --git a/src/data/two_pointing.lean b/src/data/two_pointing.lean index 3ec9c59b4b4dc..012ae3c0dfb3e 100644 --- a/src/data/two_pointing.lean +++ b/src/data/two_pointing.lean @@ -9,6 +9,10 @@ import logic.nontrivial /-! # Two-pointings +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/984 +> Any changes to this file require a corresponding PR to mathlib4. + This file defines `two_pointing α`, the type of two pointings of `α`. A two-pointing is the data of two distinct terms. diff --git a/src/logic/embedding/set.lean b/src/logic/embedding/set.lean index d7948726f6cac..d219e04aff1b2 100644 --- a/src/logic/embedding/set.lean +++ b/src/logic/embedding/set.lean @@ -9,6 +9,10 @@ import data.set.image /-! # Interactions between embeddings and sets. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/986 +> Any changes to this file require a corresponding PR to mathlib4. + -/ universes u v w x diff --git a/src/logic/equiv/embedding.lean b/src/logic/equiv/embedding.lean index 052fb9a27bf6f..547a3e582fb34 100644 --- a/src/logic/equiv/embedding.lean +++ b/src/logic/equiv/embedding.lean @@ -8,6 +8,10 @@ import logic.embedding.set /-! # Equivalences on embeddings +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/1021 +> Any changes to this file require a corresponding PR to mathlib4. + This file shows some advanced equivalences on embeddings, useful for constructing larger embeddings from smaller ones. -/ diff --git a/src/order/directed.lean b/src/order/directed.lean index 23191eacb3c45..b52108f4f7b41 100644 --- a/src/order/directed.lean +++ b/src/order/directed.lean @@ -10,6 +10,10 @@ import order.max /-! # Directed indexed families and sets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/963 +> Any changes to this file require a corresponding PR to mathlib4. + This file defines directed indexed families and directed sets. An indexed family/set is directed iff each pair of elements has a shared upper bound. diff --git a/src/order/rel_iso/set.lean b/src/order/rel_iso/set.lean index 5f236405bc97e..91a8430539550 100644 --- a/src/order/rel_iso/set.lean +++ b/src/order/rel_iso/set.lean @@ -9,6 +9,10 @@ import logic.embedding.set /-! # Interactions between relation homomorphisms and sets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/1005 +> Any changes to this file require a corresponding PR to mathlib4. + It is likely that there are better homes for many of these statement, in files further down the import graph. -/ diff --git a/src/order/well_founded.lean b/src/order/well_founded.lean index 11410254be56d..101742ceaa75f 100644 --- a/src/order/well_founded.lean +++ b/src/order/well_founded.lean @@ -9,6 +9,10 @@ import data.set.image /-! # Well-founded relations +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/970 +> Any changes to this file require a corresponding PR to mathlib4. + A relation is well-founded if it can be used for induction: for each `x`, `(∀ y, r y x → P y) → P x` implies `P x`. Well-founded relations can be used for induction and recursion, including construction of fixed points in the space of dependent functions `Π x : α , β x`. diff --git a/src/ring_theory/coprime/basic.lean b/src/ring_theory/coprime/basic.lean index 18db408295888..940441dcd4fec 100644 --- a/src/ring_theory/coprime/basic.lean +++ b/src/ring_theory/coprime/basic.lean @@ -9,6 +9,10 @@ import group_theory.group_action.units /-! # Coprime elements of a ring +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/899 +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions * `is_coprime x y`: that `x` and `y` are coprime, defined to be the existence of `a` and `b` such From 4237da0f7f6f5ddc5c2dce7faa3eb350250758b8 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 15 Dec 2022 09:57:15 +0000 Subject: [PATCH 0052/1291] =?UTF-8?q?feat(logic/lemmas):=20`(=E2=88=80=20p?= =?UTF-8?q?,=20f=20p)=20=E2=86=94=20f=20true=20=E2=88=A7=20f=20false`=20(#?= =?UTF-8?q?17944)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Expand quantifiers over `Prop` --- src/logic/lemmas.lean | 8 ++++++++ 1 file changed, 8 insertions(+) diff --git a/src/logic/lemmas.lean b/src/logic/lemmas.lean index 8464431f686cf..f868c807a62e3 100644 --- a/src/logic/lemmas.lean +++ b/src/logic/lemmas.lean @@ -3,7 +3,9 @@ Copyright (c) 2022 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ +import tactic.congr import tactic.protected +import tactic.rcases import tactic.split_ifs import logic.basic @@ -63,3 +65,9 @@ dite_dite_distrib_left lemma ite_ite_distrib_right : ite p (ite q a b) c = ite q (ite p a c) (ite p b c) := dite_dite_distrib_right + +lemma Prop.forall {f : Prop → Prop} : (∀ p, f p) ↔ f true ∧ f false := +⟨λ h, ⟨h _, h _⟩, by { rintro ⟨h₁, h₀⟩ p, by_cases hp : p; simp only [hp]; assumption }⟩ + +lemma Prop.exists {f : Prop → Prop} : (∃ p, f p) ↔ f true ∨ f false := +⟨λ ⟨p, h⟩, by refine (em p).imp _ _; intro H; convert h; simp [H], by rintro (h | h); exact ⟨_, h⟩⟩ From 633c32eab5e20f40f93323808ab83b731ee377cd Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 15 Dec 2022 11:40:49 +0000 Subject: [PATCH 0053/1291] feat(order/basic): Checking associativity in partial order (#17942) Simpler conditions to check commutativity/associativity of a function valued in a partial order. --- src/order/basic.lean | 17 +++++++++++++++++ 1 file changed, 17 insertions(+) diff --git a/src/order/basic.lean b/src/order/basic.lean index 68736eac7514b..859289add65e5 100644 --- a/src/order/basic.lean +++ b/src/order/basic.lean @@ -378,6 +378,23 @@ lemma le_implies_le_of_le_of_le {a b c d : α} [preorder α] (hca : c ≤ a) (hb a ≤ b → c ≤ d := λ hab, (hca.trans hab).trans hbd +section partial_order +variables [partial_order α] + +/-- To prove commutativity of a binary operation `○`, we only to check `a ○ b ≤ b ○ a` for all `a`, +`b`. -/ +lemma commutative_of_le {f : β → β → α} (comm : ∀ a b, f a b ≤ f b a) : ∀ a b, f a b = f b a := +λ a b, (comm _ _).antisymm $ comm _ _ + +/-- To prove associativity of a commutative binary operation `○`, we only to check +`(a ○ b) ○ c ≤ a ○ (b ○ c)` for all `a`, `b`, `c`. -/ +lemma associative_of_commutative_of_le {f : α → α → α} (comm : commutative f) + (assoc : ∀ a b c, f (f a b) c ≤ f a (f b c)) : + associative f := +λ a b c, le_antisymm (assoc _ _ _) $ by { rw [comm, comm b, comm _ c, comm a], exact assoc _ _ _ } + +end partial_order + @[ext] lemma preorder.to_has_le_injective {α : Type*} : function.injective (@preorder.to_has_le α) := From 33daf353011f93c9034730b5abcede6e37c3c40f Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 15 Dec 2022 11:40:51 +0000 Subject: [PATCH 0054/1291] =?UTF-8?q?feat(data/finset/basic):=20`finset=20?= =?UTF-8?q?=CE=B1`=20is=20well-founded=20(#17950)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit and `multiset α` too. --- src/data/finset/basic.lean | 15 +++++++++++++-- src/data/multiset/basic.lean | 2 ++ 2 files changed, 15 insertions(+), 2 deletions(-) diff --git a/src/data/finset/basic.lean b/src/data/finset/basic.lean index 2c020aac5b057..12e040fb99cf4 100644 --- a/src/data/finset/basic.lean +++ b/src/data/finset/basic.lean @@ -309,6 +309,11 @@ lemma exists_of_ssubset {s₁ s₂ : finset α} (h : s₁ ⊂ s₂) : ∃ x ∈ s₂, x ∉ s₁ := set.exists_of_ssubset h +instance is_well_founded_ssubset : is_well_founded (finset α) (⊂) := +subrelation.is_well_founded (inv_image _ _) $ λ _ _, val_lt_iff.2 + +instance is_well_founded_lt : well_founded_lt (finset α) := finset.is_well_founded_ssubset + end subset -- TODO: these should be global attributes, but this will require fixing other files @@ -2007,7 +2012,7 @@ def not_mem_range_equiv (k : ℕ) : {n // n ∉ range k} ≃ ℕ := /-! ### dedup on list and multiset -/ namespace multiset -variable [decidable_eq α] +variables [decidable_eq α] {s t : multiset α} /-- `to_finset s` removes duplicates from the multiset `s` to produce a finset. -/ def to_finset (s : multiset α) : finset α := ⟨_, nodup_dedup s⟩ @@ -2054,9 +2059,12 @@ by ext; simp @[simp] theorem to_finset_eq_empty {m : multiset α} : m.to_finset = ∅ ↔ m = 0 := finset.val_inj.symm.trans multiset.dedup_eq_zero -@[simp] lemma to_finset_subset (s t : multiset α) : s.to_finset ⊆ t.to_finset ↔ s ⊆ t := +@[simp] lemma to_finset_subset : s.to_finset ⊆ t.to_finset ↔ s ⊆ t := by simp only [finset.subset_iff, multiset.subset_iff, multiset.mem_to_finset] +@[simp] lemma to_finset_ssubset : s.to_finset ⊂ t.to_finset ↔ s ⊂ t := +by { simp_rw [finset.ssubset_def, to_finset_subset], refl } + @[simp] lemma to_finset_dedup (m : multiset α) : m.dedup.to_finset = m.to_finset := by simp_rw [to_finset, dedup_idempotent] @@ -2065,6 +2073,9 @@ by simp_rw [to_finset, dedup_idempotent] (m.dedup.bind f).to_finset = (m.bind f).to_finset := by simp_rw [to_finset, dedup_bind_dedup] +instance is_well_founded_ssubset : is_well_founded (multiset β) (⊂) := +subrelation.is_well_founded (inv_image _ _) $ λ _ _, by classical; exact to_finset_ssubset.2 + end multiset namespace finset diff --git a/src/data/multiset/basic.lean b/src/data/multiset/basic.lean index 20c91d032a450..1881abd1b3170 100644 --- a/src/data/multiset/basic.lean +++ b/src/data/multiset/basic.lean @@ -607,6 +607,8 @@ by { dunfold strong_downward_induction_on, rw strong_downward_induction } lemma well_founded_lt : well_founded ((<) : multiset α → multiset α → Prop) := subrelation.wf (λ _ _, multiset.card_lt_of_lt) (measure_wf multiset.card) +instance is_well_founded_lt : _root_.well_founded_lt (multiset α) := ⟨well_founded_lt⟩ + /-! ### `multiset.repeat` -/ /-- `repeat a n` is the multiset containing only `a` with multiplicity `n`. -/ From 49420fb14629125899f2b4831fca875ccb9411ad Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Thu, 15 Dec 2022 11:40:52 +0000 Subject: [PATCH 0055/1291] chore(number_theory/arithmetic_function): golf (#17953) Also generalize `coe_zeta_smul_apply`. --- src/number_theory/arithmetic_function.lean | 73 +++++----------------- 1 file changed, 14 insertions(+), 59 deletions(-) diff --git a/src/number_theory/arithmetic_function.lean b/src/number_theory/arithmetic_function.lean index 30b0371e27668..2d05ecfe97d50 100644 --- a/src/number_theory/arithmetic_function.lean +++ b/src/number_theory/arithmetic_function.lean @@ -337,78 +337,33 @@ lemma zeta_apply {x : ℕ} : ζ x = if (x = 0) then 0 else 1 := rfl lemma zeta_apply_ne {x : ℕ} (h : x ≠ 0) : ζ x = 1 := if_neg h -@[simp] -theorem coe_zeta_mul_apply [semiring R] {f : arithmetic_function R} {x : ℕ} : - (↑ζ * f) x = ∑ i in divisors x, f i := -begin - rw mul_apply, - transitivity ∑ i in divisors_antidiagonal x, f i.snd, - { apply sum_congr rfl, - intros i hi, - rcases mem_divisors_antidiagonal.1 hi with ⟨rfl, h⟩, - rw [nat_coe_apply, zeta_apply_ne (left_ne_zero_of_mul h), cast_one, one_mul] }, - { apply sum_bij (λ i h, prod.snd i), - { rintros ⟨a, b⟩ h, simp [snd_mem_divisors_of_mem_antidiagonal h] }, - { rintros ⟨a, b⟩ h, refl }, - { rintros ⟨a1, b1⟩ ⟨a2, b2⟩ h1 h2 h, - dsimp at h, - rw h at *, - rw mem_divisors_antidiagonal at *, - ext, swap, {refl}, - simp only [prod.fst, prod.snd] at *, - apply nat.eq_of_mul_eq_mul_right _ (eq.trans h1.1 h2.1.symm), - rcases h1 with ⟨rfl, h⟩, - apply nat.pos_of_ne_zero (right_ne_zero_of_mul h) }, - { intros a ha, - rcases mem_divisors.1 ha with ⟨⟨b, rfl⟩, ne0⟩, - use (b, a), - simp [ne0, mul_comm] } } -end - -theorem coe_zeta_smul_apply {M : Type*} [comm_ring R] [add_comm_group M] [module R M] +@[simp] theorem coe_zeta_smul_apply {M} [semiring R] [add_comm_monoid M] [module R M] {f : arithmetic_function M} {x : ℕ} : ((↑ζ : arithmetic_function R) • f) x = ∑ i in divisors x, f i := begin rw smul_apply, transitivity ∑ i in divisors_antidiagonal x, f i.snd, - { apply sum_congr rfl, - intros i hi, + { refine sum_congr rfl (λ i hi, _), rcases mem_divisors_antidiagonal.1 hi with ⟨rfl, h⟩, rw [nat_coe_apply, zeta_apply_ne (left_ne_zero_of_mul h), cast_one, one_smul] }, - { apply sum_bij (λ i h, prod.snd i), - { rintros ⟨a, b⟩ h, simp [snd_mem_divisors_of_mem_antidiagonal h] }, - { rintros ⟨a, b⟩ h, refl }, - { rintros ⟨a1, b1⟩ ⟨a2, b2⟩ h1 h2 h, - dsimp at h, - rw h at *, - rw mem_divisors_antidiagonal at *, - ext, swap, {refl}, - simp only [prod.fst, prod.snd] at *, - apply nat.eq_of_mul_eq_mul_right _ (eq.trans h1.1 h2.1.symm), - rcases h1 with ⟨rfl, h⟩, - apply nat.pos_of_ne_zero (right_ne_zero_of_mul h) }, - { intros a ha, - rcases mem_divisors.1 ha with ⟨⟨b, rfl⟩, ne0⟩, - use (b, a), - simp [ne0, mul_comm] } } + { rw [← map_div_left_divisors, sum_map, function.embedding.coe_fn_mk] } end +@[simp] +theorem coe_zeta_mul_apply [semiring R] {f : arithmetic_function R} {x : ℕ} : + (↑ζ * f) x = ∑ i in divisors x, f i := +coe_zeta_smul_apply + @[simp] theorem coe_mul_zeta_apply [semiring R] {f : arithmetic_function R} {x : ℕ} : (f * ζ) x = ∑ i in divisors x, f i := begin - apply mul_opposite.op_injective, - rw [op_sum], - convert @coe_zeta_mul_apply Rᵐᵒᵖ _ { to_fun := mul_opposite.op ∘ f, map_zero' := by simp} x, - rw [mul_apply, mul_apply, op_sum], - conv_lhs { rw ← map_swap_divisors_antidiagonal, }, - rw sum_map, - apply sum_congr rfl, - intros y hy, - by_cases h1 : y.fst = 0, - { simp [function.comp_apply, h1] }, - { simp only [h1, mul_one, one_mul, prod.fst_swap, function.embedding.coe_fn_mk, prod.snd_swap, - if_false, zeta_apply, zero_hom.coe_mk, nat_coe_apply, cast_one] } + rw mul_apply, + transitivity ∑ i in divisors_antidiagonal x, f i.1, + { refine sum_congr rfl (λ i hi, _), + rcases mem_divisors_antidiagonal.1 hi with ⟨rfl, h⟩, + rw [nat_coe_apply, zeta_apply_ne (right_ne_zero_of_mul h), cast_one, mul_one] }, + { rw [← map_div_right_divisors, sum_map, function.embedding.coe_fn_mk] } end theorem zeta_mul_apply {f : arithmetic_function ℕ} {x : ℕ} : From 616e053df08ef2addfe9ba9ca1d075785d252562 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 15 Dec 2022 14:52:57 +0000 Subject: [PATCH 0056/1291] chore(data/set/image): Turn lemma around (#17945) Turn `set.image_inter` around to match all other distributivity lemmas (`set.image_union`, `set.image_Inter`, `set.image_Union`, `finset.image_inter`, `finset.image_union`, ...) --- src/algebraic_geometry/ringed_space.lean | 2 +- src/data/set/image.lean | 10 +++++----- src/data/set/n_ary.lean | 4 ++-- src/data/set/pointwise/smul.lean | 2 +- src/measure_theory/measure/measure_space.lean | 6 ++---- src/order/filter/basic.lean | 2 +- src/order/upper_lower.lean | 6 ++---- src/topology/quasi_separated.lean | 2 +- 8 files changed, 15 insertions(+), 19 deletions(-) diff --git a/src/algebraic_geometry/ringed_space.lean b/src/algebraic_geometry/ringed_space.lean index 919792ac6b328..506bc9d9cdbb6 100644 --- a/src/algebraic_geometry/ringed_space.lean +++ b/src/algebraic_geometry/ringed_space.lean @@ -174,7 +174,7 @@ end begin ext1, dsimp [RingedSpace.basic_open], - rw set.image_inter subtype.coe_injective, + rw ←set.image_inter subtype.coe_injective, congr, ext, simp_rw map_mul, diff --git a/src/data/set/image.lean b/src/data/set/image.lean index 7b49083f2bbf5..0a6463c1cb78d 100644 --- a/src/data/set/image.lean +++ b/src/data/set/image.lean @@ -229,15 +229,15 @@ lemma image_inter_subset (f : α → β) (s t : set α) : subset_inter (image_subset _ $ inter_subset_left _ _) (image_subset _ $ inter_subset_right _ _) theorem image_inter_on {f : α → β} {s t : set α} (h : ∀x∈t, ∀y∈s, f x = f y → x = y) : - f '' s ∩ f '' t = f '' (s ∩ t) := -subset.antisymm + f '' (s ∩ t) = f '' s ∩ f '' t := +(image_inter_subset _ _ _).antisymm (assume b ⟨⟨a₁, ha₁, h₁⟩, ⟨a₂, ha₂, h₂⟩⟩, have a₂ = a₁, from h _ ha₂ _ ha₁ (by simp *), ⟨a₁, ⟨ha₁, this ▸ ha₂⟩, h₁⟩) - (image_inter_subset _ _ _) + theorem image_inter {f : α → β} {s t : set α} (H : injective f) : - f '' s ∩ f '' t = f '' (s ∩ t) := + f '' (s ∩ t) = f '' s ∩ f '' t := image_inter_on (assume x _ y _ h, H h) theorem image_univ_of_surjective {ι : Type*} {f : ι → β} (H : surjective f) : f '' univ = univ := @@ -300,7 +300,7 @@ theorem mem_image_iff_of_inverse {f : α → β} {g : β → α} {b : β} {s : s by rw image_eq_preimage_of_inverse h₁ h₂; refl theorem image_compl_subset {f : α → β} {s : set α} (H : injective f) : f '' sᶜ ⊆ (f '' s)ᶜ := -disjoint.subset_compl_left $ by simp [disjoint_iff_inf_le, image_inter H] +disjoint.subset_compl_left $ by simp [disjoint_iff_inf_le, ←image_inter H] theorem subset_image_compl {f : α → β} {s : set α} (H : surjective f) : (f '' s)ᶜ ⊆ f '' sᶜ := compl_subset_iff_union.2 $ diff --git a/src/data/set/n_ary.lean b/src/data/set/n_ary.lean index 6eae6cc3b6262..5c43a8082afff 100644 --- a/src/data/set/n_ary.lean +++ b/src/data/set/n_ary.lean @@ -104,10 +104,10 @@ begin end lemma image2_inter_left (hf : injective2 f) : image2 f (s ∩ s') t = image2 f s t ∩ image2 f s' t := -by simp_rw [←image_uncurry_prod, inter_prod, ←image_inter hf.uncurry] +by simp_rw [←image_uncurry_prod, inter_prod, image_inter hf.uncurry] lemma image2_inter_right (hf : injective2 f) : image2 f s (t ∩ t') = image2 f s t ∩ image2 f s t' := -by simp_rw [←image_uncurry_prod, prod_inter, ←image_inter hf.uncurry] +by simp_rw [←image_uncurry_prod, prod_inter, image_inter hf.uncurry] @[simp] lemma image2_empty_left : image2 f ∅ t = ∅ := ext $ by simp @[simp] lemma image2_empty_right : image2 f s ∅ = ∅ := ext $ by simp diff --git a/src/data/set/pointwise/smul.lean b/src/data/set/pointwise/smul.lean index 77892d5c43e5a..dd9a12dad6c23 100644 --- a/src/data/set/pointwise/smul.lean +++ b/src/data/set/pointwise/smul.lean @@ -473,7 +473,7 @@ iff.symm $ (image_subset_iff).trans $ iff.symm $ iff_of_eq $ congr_arg _ $ image_equiv_eq_preimage_symm _ $ mul_action.to_perm _ @[to_additive] lemma smul_set_inter : a • (s ∩ t) = a • s ∩ a • t := -(image_inter $ mul_action.injective a).symm +image_inter $ mul_action.injective a @[to_additive] lemma smul_set_sdiff : a • (s \ t) = a • s \ a • t := image_diff (mul_action.injective a) _ _ diff --git a/src/measure_theory/measure/measure_space.lean b/src/measure_theory/measure/measure_space.lean index 3401782f58569..68716ac6dd571 100644 --- a/src/measure_theory/measure/measure_space.lean +++ b/src/measure_theory/measure/measure_space.lean @@ -1089,8 +1089,7 @@ def comapₗ [measurable_space α] (f : α → β) : measure β →ₗ[ℝ≥0 if hf : injective f ∧ ∀ s, measurable_set s → measurable_set (f '' s) then lift_linear (outer_measure.comap f) $ λ μ s hs t, begin - simp only [coe_to_outer_measure, outer_measure.comap_apply, ← image_inter hf.1, - image_diff hf.1], + simp only [coe_to_outer_measure, outer_measure.comap_apply, image_inter hf.1, image_diff hf.1], apply le_to_outer_measure_caratheodory, exact hf.2 s hs end @@ -1111,8 +1110,7 @@ def comap [measurable_space α] (f : α → β) (μ : measure β) : measure α : if hf : injective f ∧ ∀ s, measurable_set s → null_measurable_set (f '' s) μ then (outer_measure.comap f μ.to_outer_measure).to_measure $ λ s hs t, begin - simp only [coe_to_outer_measure, outer_measure.comap_apply, ← image_inter hf.1, - image_diff hf.1], + simp only [coe_to_outer_measure, outer_measure.comap_apply, image_inter hf.1, image_diff hf.1], exact (measure_inter_add_diff₀ _ (hf.2 s hs)).symm end else 0 diff --git a/src/order/filter/basic.lean b/src/order/filter/basic.lean index 607653b6c458b..125aef03200b7 100644 --- a/src/order/filter/basic.lean +++ b/src/order/filter/basic.lean @@ -2080,7 +2080,7 @@ begin refine map_inf_le.antisymm _, rintro t ⟨s₁, hs₁, s₂, hs₂, ht : m ⁻¹' t = s₁ ∩ s₂⟩, refine mem_inf_of_inter (image_mem_map hs₁) (image_mem_map hs₂) _, - rw [image_inter h, image_subset_iff, ht] + rw [←image_inter h, image_subset_iff, ht] end lemma map_inf' {f g : filter α} {m : α → β} {t : set α} (htf : t ∈ f) (htg : t ∈ g) diff --git a/src/order/upper_lower.lean b/src/order/upper_lower.lean index 0864b562e6a5b..61b3c5a6c2b40 100644 --- a/src/order/upper_lower.lean +++ b/src/order/upper_lower.lean @@ -548,8 +548,7 @@ variables (f s t) @[simp, norm_cast] lemma coe_map : (map f s : set β) = f '' s := rfl -@[simp] protected lemma map_sup : map f (s ⊔ t) = map f s ⊔ map f t := -ext $ (image_inter f.injective).symm +@[simp] protected lemma map_sup : map f (s ⊔ t) = map f s ⊔ map f t := ext $ image_inter f.injective @[simp] protected lemma map_inf : map f (s ⊓ t) = map f s ⊓ map f t := ext $ image_union _ _ _ @[simp] protected lemma map_top : map f ⊤ = ⊤ := ext $ image_empty _ @[simp] protected lemma map_bot : map f ⊥ = ⊥ := ext $ image_univ_of_surjective f.surjective @@ -594,8 +593,7 @@ variables (f s t) @[simp, norm_cast] lemma coe_map : (map f s : set β) = f '' s := rfl @[simp] protected lemma map_sup : map f (s ⊔ t) = map f s ⊔ map f t := ext $ image_union _ _ _ -@[simp] protected lemma map_inf : map f (s ⊓ t) = map f s ⊓ map f t := -ext $ (image_inter f.injective).symm +@[simp] protected lemma map_inf : map f (s ⊓ t) = map f s ⊓ map f t := ext $ image_inter f.injective @[simp] protected lemma map_top : map f ⊤ = ⊤ := ext $ image_univ_of_surjective f.surjective @[simp] protected lemma map_bot : map f ⊥ = ⊥ := ext $ image_empty _ @[simp] protected lemma map_Sup (S : set (lower_set α)) : map f (Sup S) = ⨆ s ∈ S, map f s := diff --git a/src/topology/quasi_separated.lean b/src/topology/quasi_separated.lean index 93b8a30bdbfb9..b4ad41e647c87 100644 --- a/src/topology/quasi_separated.lean +++ b/src/topology/quasi_separated.lean @@ -83,7 +83,7 @@ lemma open_embedding.is_quasi_separated_iff (h : open_embedding f) {s : set α} begin refine ⟨λ hs, hs.image_of_embedding h.to_embedding, _⟩, intros H U V hU hU' hU'' hV hV' hV'', - rw [h.to_embedding.is_compact_iff_is_compact_image, ← set.image_inter h.inj], + rw [h.to_embedding.is_compact_iff_is_compact_image, set.image_inter h.inj], exact H (f '' U) (f '' V) (set.image_subset _ hU) (h.is_open_map _ hU') (hU''.image h.continuous) (set.image_subset _ hV) (h.is_open_map _ hV') (hV''.image h.continuous) From ac3910d7f80f7d2d8abaf73af3ba751f60e2d009 Mon Sep 17 00:00:00 2001 From: Kevin Buzzard Date: Thu, 15 Dec 2022 14:52:58 +0000 Subject: [PATCH 0057/1291] fix(algebra/opposites): fix types in unop_div (#17955) This lemma only typechecked because of defeq abuse, which was presumably not intended. Remark: refactoring `mul_opposite` to be a structure in Lean 4 was what picked up on this error. --- src/algebra/opposites.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/algebra/opposites.lean b/src/algebra/opposites.lean index e4d3e2d8e1314..cbe829d2c6471 100644 --- a/src/algebra/opposites.lean +++ b/src/algebra/opposites.lean @@ -210,6 +210,6 @@ instance [has_involutive_inv α] : has_involutive_inv αᵃᵒᵖ := instance [has_div α] : has_div αᵃᵒᵖ := { div := λ a b, op (unop a / unop b) } @[simp] lemma op_div [has_div α] (a b : α) : op (a / b) = op a / op b := rfl -@[simp] lemma unop_div [has_div α] (a b : α) : unop (a / b) = unop a / unop b := rfl +@[simp] lemma unop_div [has_div α] (a b : αᵃᵒᵖ) : unop (a / b) = unop a / unop b := rfl end add_opposite From e3a8dff0bb416efaa28a590d82bafdbc52e4fbb4 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9mi=20Bottinelli?= Date: Thu, 15 Dec 2022 18:06:17 +0000 Subject: [PATCH 0058/1291] feat(combinatorics/simple_graph/*): some more `from_edge_set` and `edge_set` lemmas (#17723) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Kyle Miller Co-authored-by: Yaël Dillies --- src/combinatorics/simple_graph/basic.lean | 154 +++++++++++++++--- .../simple_graph/connectivity.lean | 15 +- src/data/fintype/basic.lean | 8 +- 3 files changed, 145 insertions(+), 32 deletions(-) diff --git a/src/combinatorics/simple_graph/basic.lean b/src/combinatorics/simple_graph/basic.lean index 79a8aff3c977b..593e33e9fb4b8 100644 --- a/src/combinatorics/simple_graph/basic.lean +++ b/src/combinatorics/simple_graph/basic.lean @@ -73,7 +73,9 @@ finitely many vertices. in order to start learning what the combinatorics hierarchy should look like. -/ -open finset + +open finset function + universes u v w /-- @@ -270,6 +272,9 @@ def neighbor_set (v : V) : set V := set_of (G.adj v) instance neighbor_set.mem_decidable (v : V) [decidable_rel G.adj] : decidable_pred (∈ G.neighbor_set v) := by { unfold neighbor_set, apply_instance } +section edge_set +variables {G₁ G₂ : simple_graph V} + /-- The edges of G consist of the unordered pairs of vertices related by `G.adj`. @@ -277,12 +282,56 @@ The edges of G consist of the unordered pairs of vertices related by The way `edge_set` is defined is such that `mem_edge_set` is proved by `refl`. (That is, `⟦(v, w)⟧ ∈ G.edge_set` is definitionally equal to `G.adj v w`.) -/ -def edge_set : set (sym2 V) := sym2.from_rel G.symm +def edge_set : simple_graph V ↪o set (sym2 V) := +order_embedding.of_map_le_iff (λ G, sym2.from_rel G.symm) $ + λ G G', ⟨λ h a b, @h ⟦(a, b)⟧, λ h e, sym2.ind @h e⟩ @[simp] lemma mem_edge_set : ⟦(v, w)⟧ ∈ G.edge_set ↔ G.adj v w := iff.rfl -lemma edge_set_mono {G G' : simple_graph V} (h : G ≤ G') : G.edge_set ⊆ G'.edge_set := -λ e, sym2.ind (λ v w, @h v w) e +lemma not_is_diag_of_mem_edge_set : e ∈ G.edge_set → ¬ e.is_diag := sym2.ind (λ v w, adj.ne) e + +@[simp] lemma edge_set_inj : G₁.edge_set = G₂.edge_set ↔ G₁ = G₂ := +(edge_set : simple_graph V ↪o set (sym2 V)).eq_iff_eq + +@[simp] lemma edge_set_subset_edge_set : G₁.edge_set ⊆ G₂.edge_set ↔ G₁ ≤ G₂ := +(edge_set : simple_graph V ↪o set (sym2 V)).le_iff_le + +@[simp] lemma edge_set_ssubset_edge_set : G₁.edge_set ⊂ G₂.edge_set ↔ G₁ < G₂ := +(edge_set : simple_graph V ↪o set (sym2 V)).lt_iff_lt + +lemma edge_set_injective : injective (edge_set : simple_graph V → set (sym2 V)) := +edge_set.injective + +alias edge_set_subset_edge_set ↔ _ edge_set_mono +alias edge_set_ssubset_edge_set ↔ _ edge_set_strict_mono + +attribute [mono] edge_set_mono edge_set_strict_mono + +variables (G₁ G₂) + +@[simp] lemma edge_set_bot : (⊥ : simple_graph V).edge_set = ∅ := sym2.from_rel_bot + +@[simp] lemma edge_set_sup : (G₁ ⊔ G₂).edge_set = G₁.edge_set ∪ G₂.edge_set := +by { ext ⟨x, y⟩, refl } + +@[simp] lemma edge_set_inf : (G₁ ⊓ G₂).edge_set = G₁.edge_set ∩ G₂.edge_set := +by { ext ⟨x, y⟩, refl } + +@[simp] lemma edge_set_sdiff : (G₁ \ G₂).edge_set = G₁.edge_set \ G₂.edge_set := +by { ext ⟨x, y⟩, refl } + +/-- +This lemma, combined with `edge_set_sdiff` and `edge_set_from_edge_set`, +allows proving `(G \ from_edge_set s).edge_set = G.edge_set \ s` by `simp`. +-/ +@[simp] lemma edge_set_sdiff_sdiff_is_diag (G : simple_graph V) (s : set (sym2 V)) : + G.edge_set \ (s \ {e | e.is_diag}) = G.edge_set \ s := +begin + ext e, + simp only [set.mem_diff, set.mem_set_of_eq, not_and, not_not, and.congr_right_iff], + intro h, + simp only [G.not_is_diag_of_mem_edge_set h, imp_false], +end /-- Two vertices are adjacent iff there is an edge between them. The @@ -314,9 +363,26 @@ end instance decidable_mem_edge_set [decidable_rel G.adj] : decidable_pred (∈ G.edge_set) := sym2.from_rel.decidable_pred _ -instance edges_fintype [decidable_eq V] [fintype V] [decidable_rel G.adj] : +instance fintype_edge_set [decidable_eq V] [fintype V] [decidable_rel G.adj] : fintype G.edge_set := subtype.fintype _ +instance fintype_edge_set_bot : fintype (⊥ : simple_graph V).edge_set := +by { rw edge_set_bot, apply_instance } + +instance fintype_edge_set_sup [decidable_eq V] [fintype G₁.edge_set] [fintype G₂.edge_set] : + fintype (G₁ ⊔ G₂).edge_set := +by { rw edge_set_sup, apply_instance } + +instance fintype_edge_set_inf [decidable_eq V] [fintype G₁.edge_set] [fintype G₂.edge_set] : + fintype (G₁ ⊓ G₂).edge_set := +by { rw edge_set_inf, exact set.fintype_inter _ _ } + +instance fintype_edge_set_sdiff [decidable_eq V] [fintype G₁.edge_set] [fintype G₂.edge_set] : + fintype (G₁ \ G₂).edge_set := +by { rw edge_set_sdiff, exact set.fintype_diff _ _ } + +end edge_set + section from_edge_set variable (s : set (sym2 V)) @@ -332,7 +398,7 @@ def from_edge_set : simple_graph V := -- Note: we need to make sure `from_edge_set_adj` and this lemma are confluent. -- In particular, both yield `⟦(u, v)⟧ ∈ (from_edge_set s).edge_set` ==> `⟦(v, w)⟧ ∈ s ∧ v ≠ w`. -@[simp] lemma edge_set_from_edge_set : (from_edge_set s).edge_set = {e ∈ s | ¬ e.is_diag} := +@[simp] lemma edge_set_from_edge_set : (from_edge_set s).edge_set = s \ {e | e.is_diag} := by { ext e, exact sym2.ind (by simp) e } @[simp] lemma from_edge_set_edge_set : from_edge_set G.edge_set = G := @@ -344,6 +410,26 @@ by { ext v w, simp only [from_edge_set_adj, set.mem_empty_iff_false, false_and, @[simp] lemma from_edge_set_univ : from_edge_set (set.univ : set (sym2 V)) = ⊤ := by { ext v w, simp only [from_edge_set_adj, set.mem_univ, true_and, top_adj] } +@[simp] lemma from_edge_set_inf (s t : set (sym2 V)) : + from_edge_set s ⊓ from_edge_set t = from_edge_set (s ∩ t) := +by { ext v w, simp only [from_edge_set_adj, set.mem_inter_iff, ne.def, inf_adj], tauto, } + +@[simp] lemma from_edge_set_sup (s t : set (sym2 V)) : + from_edge_set s ⊔ from_edge_set t = from_edge_set (s ∪ t) := +by { ext v w, simp [set.mem_union, or_and_distrib_right], } + +@[simp] lemma from_edge_set_sdiff (s t : set (sym2 V)) : + from_edge_set s \ from_edge_set t = from_edge_set (s \ t) := +by { ext v w, split; simp { contextual := tt }, } + +@[mono] +lemma from_edge_set_mono {s t : set (sym2 V)} (h : s ⊆ t) : from_edge_set s ≤ from_edge_set t := +begin + rintro v w, + simp only [from_edge_set_adj, ne.def, not_false_iff, and_true, and_imp] {contextual := tt}, + exact λ vws _, h vws, +end + instance [decidable_eq V] [fintype s] : fintype (from_edge_set s).edge_set := by { rw edge_set_from_edge_set s, apply_instance } @@ -487,30 +573,54 @@ end instance decidable_mem_incidence_set [decidable_eq V] [decidable_rel G.adj] (v : V) : decidable_pred (∈ G.incidence_set v) := λ e, and.decidable +section edge_finset +variables {G₁ G₂ : simple_graph V} [fintype G.edge_set] [fintype G₁.edge_set] [fintype G₂.edge_set] + /-- The `edge_set` of the graph as a `finset`. -/ -@[reducible] def edge_finset [fintype G.edge_set] : finset (sym2 V) := -set.to_finset G.edge_set - -@[simp] lemma mem_edge_finset [fintype G.edge_set] (e : sym2 V) : - e ∈ G.edge_finset ↔ e ∈ G.edge_set := -set.mem_to_finset +@[reducible] def edge_finset : finset (sym2 V) := set.to_finset G.edge_set -@[simp, norm_cast] lemma coe_edge_finset [fintype G.edge_set] : - (G.edge_finset : set (sym2 V)) = G.edge_set := +@[simp, norm_cast] lemma coe_edge_finset : (G.edge_finset : set (sym2 V)) = G.edge_set := set.coe_to_finset _ -lemma edge_finset_mono {G G' : simple_graph V} [fintype G.edge_set] [fintype G'.edge_set] : - G ≤ G' → G.edge_finset ⊆ G'.edge_finset := -by { simp_rw [←coe_subset, coe_edge_finset], exact edge_set_mono } +variables {G} + +@[simp] lemma mem_edge_finset : e ∈ G.edge_finset ↔ e ∈ G.edge_set := set.mem_to_finset + +lemma not_is_diag_of_mem_edge_finset : e ∈ G.edge_finset → ¬ e.is_diag := +not_is_diag_of_mem_edge_set _ ∘ mem_edge_finset.1 + +@[simp] lemma edge_finset_inj : G₁.edge_finset = G₂.edge_finset ↔ G₁ = G₂ := by simp [edge_finset] + +@[simp] lemma edge_finset_subset_edge_finset : G₁.edge_finset ⊆ G₂.edge_finset ↔ G₁ ≤ G₂ := +by simp [edge_finset] + +@[simp] lemma edge_finset_ssubset_edge_finset : G₁.edge_finset ⊂ G₂.edge_finset ↔ G₁ < G₂ := +by simp [edge_finset] + +alias edge_finset_subset_edge_finset ↔ _ edge_finset_mono +alias edge_finset_ssubset_edge_finset ↔ _ edge_finset_strict_mono + +attribute [mono] edge_finset_mono edge_finset_strict_mono + +@[simp] lemma edge_finset_bot : (⊥ : simple_graph V).edge_finset = ∅ := by simp [edge_finset] + +@[simp] lemma edge_finset_sup : (G₁ ⊔ G₂).edge_finset = G₁.edge_finset ∪ G₂.edge_finset := +by simp [edge_finset] + +@[simp] lemma edge_finset_inf : (G₁ ⊓ G₂).edge_finset = G₁.edge_finset ∩ G₂.edge_finset := +by simp [edge_finset] + +@[simp] lemma edge_finset_sdiff : (G₁ \ G₂).edge_finset = G₁.edge_finset \ G₂.edge_finset := +by simp [edge_finset] + +lemma edge_finset_card : G.edge_finset.card = fintype.card G.edge_set := set.to_finset_card _ -lemma edge_finset_card [fintype G.edge_set] : G.edge_finset.card = fintype.card G.edge_set := -set.to_finset_card _ +@[simp] lemma edge_set_univ_card : (univ : finset G.edge_set).card = G.edge_finset.card := +fintype.card_of_subtype G.edge_finset $ λ _, mem_edge_finset -@[simp] lemma edge_set_univ_card [fintype G.edge_set] : - (univ : finset G.edge_set).card = G.edge_finset.card := -fintype.card_of_subtype G.edge_finset (mem_edge_finset _) +end edge_finset @[simp] lemma mem_neighbor_set (v w : V) : w ∈ G.neighbor_set v ↔ G.adj v w := iff.rfl diff --git a/src/combinatorics/simple_graph/connectivity.lean b/src/combinatorics/simple_graph/connectivity.lean index 1fd28aca8c8cc..84f21619b2cdb 100644 --- a/src/combinatorics/simple_graph/connectivity.lean +++ b/src/combinatorics/simple_graph/connectivity.lean @@ -1651,22 +1651,25 @@ section bridge_edges are no longer reachable from one another. -/ def is_bridge (G : simple_graph V) (e : sym2 V) : Prop := e ∈ G.edge_set ∧ -sym2.lift ⟨λ v w, ¬ (G.delete_edges {e}).reachable v w, by simp [reachable_comm]⟩ e +sym2.lift ⟨λ v w, ¬ (G \ from_edge_set {e}).reachable v w, by simp [reachable_comm]⟩ e lemma is_bridge_iff {u v : V} : - G.is_bridge ⟦(u, v)⟧ ↔ G.adj u v ∧ ¬ (G.delete_edges {⟦(u, v)⟧}).reachable u v := iff.rfl + G.is_bridge ⟦(u, v)⟧ ↔ G.adj u v ∧ ¬ (G \ from_edge_set {⟦(u, v)⟧}).reachable u v := iff.rfl lemma reachable_delete_edges_iff_exists_walk {v w : V} : - (G.delete_edges {⟦(v, w)⟧}).reachable v w ↔ ∃ (p : G.walk v w), ¬ ⟦(v, w)⟧ ∈ p.edges := + (G \ from_edge_set {⟦(v, w)⟧}).reachable v w ↔ ∃ (p : G.walk v w), ¬ ⟦(v, w)⟧ ∈ p.edges := begin split, { rintro ⟨p⟩, - use p.map (hom.map_spanning_subgraphs (G.delete_edges_le _)), + use p.map (hom.map_spanning_subgraphs (by simp)), simp_rw [walk.edges_map, list.mem_map, hom.map_spanning_subgraphs_apply, sym2.map_id', id.def], rintro ⟨e, h, rfl⟩, simpa using p.edges_subset_edge_set h, }, { rintro ⟨p, h⟩, - exact ⟨p.to_delete_edge _ h⟩, }, + refine ⟨p.transfer _ (λ e ep, _)⟩, + simp only [edge_set_sdiff, edge_set_from_edge_set, edge_set_sdiff_sdiff_is_diag, + set.mem_diff, set.mem_singleton_iff], + exact ⟨p.edges_subset_edge_set ep, λ h', h (h' ▸ ep)⟩, }, end lemma is_bridge_iff_adj_and_forall_walk_mem_edges {v w : V} : @@ -1709,7 +1712,7 @@ begin end lemma adj_and_reachable_delete_edges_iff_exists_cycle {v w : V} : - G.adj v w ∧ (G.delete_edges {⟦(v, w)⟧}).reachable v w ↔ + G.adj v w ∧ (G \ from_edge_set {⟦(v, w)⟧}).reachable v w ↔ ∃ (u : V) (p : G.walk u u), p.is_cycle ∧ ⟦(v, w)⟧ ∈ p.edges := begin classical, diff --git a/src/data/fintype/basic.lean b/src/data/fintype/basic.lean index ab0c192c6ade2..2d11dac574f2c 100644 --- a/src/data/fintype/basic.lean +++ b/src/data/fintype/basic.lean @@ -482,15 +482,15 @@ by simp only [finset.ssubset_def, to_finset_subset, ssubset_def] disjoint s.to_finset t.to_finset ↔ disjoint s t := by simp only [←disjoint_coe, coe_to_finset] -lemma to_finset_inter {α : Type*} [decidable_eq α] (s t : set α) [fintype (s ∩ t : set α)] +@[simp] lemma to_finset_inter {α : Type*} [decidable_eq α] (s t : set α) [fintype (s ∩ t : set α)] [fintype s] [fintype t] : (s ∩ t).to_finset = s.to_finset ∩ t.to_finset := by { ext, simp } -lemma to_finset_union {α : Type*} [decidable_eq α] (s t : set α) [fintype (s ∪ t : set α)] +@[simp] lemma to_finset_union {α : Type*} [decidable_eq α] (s t : set α) [fintype (s ∪ t : set α)] [fintype s] [fintype t] : (s ∪ t).to_finset = s.to_finset ∪ t.to_finset := by { ext, simp } -lemma to_finset_diff {α : Type*} [decidable_eq α] (s t : set α) [fintype s] [fintype t] +@[simp] lemma to_finset_diff {α : Type*} [decidable_eq α] (s t : set α) [fintype s] [fintype t] [fintype (s \ t : set α)] : (s \ t).to_finset = s.to_finset \ t.to_finset := by { ext, simp } @@ -498,7 +498,7 @@ lemma to_finset_ne_eq_erase {α : Type*} [decidable_eq α] [fintype α] (a : α) [fintype {x : α | x ≠ a}] : {x : α | x ≠ a}.to_finset = finset.univ.erase a := by { ext, simp } -theorem to_finset_compl [decidable_eq α] [fintype α] (s : set α) [fintype s] [fintype ↥sᶜ] : +@[simp] theorem to_finset_compl [decidable_eq α] [fintype α] (s : set α) [fintype s] [fintype ↥sᶜ] : (sᶜ).to_finset = s.to_finsetᶜ := by { ext, simp } From 302eab4f46abb63de520828de78c04cb0f9b5836 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 15 Dec 2022 18:06:19 +0000 Subject: [PATCH 0059/1291] chore(data/finset/lattice): Protect ambiguous lemmas (#17938) Protect a few `finset` lemmas that conflict with root ones when `finset` is open. --- src/data/finset/lattice.lean | 61 +++++++++---------- src/data/finset/sigma.lean | 7 +-- src/data/polynomial/degree/lemmas.lean | 2 +- .../polynomial/degree/trailing_degree.lean | 2 +- src/data/polynomial/derivative.lean | 2 +- src/data/polynomial/unit_trinomial.lean | 8 +-- 6 files changed, 37 insertions(+), 45 deletions(-) diff --git a/src/data/finset/lattice.lean b/src/data/finset/lattice.lean index f86ede26b6035..8ba891c9b7c67 100644 --- a/src/data/finset/lattice.lean +++ b/src/data/finset/lattice.lean @@ -72,6 +72,12 @@ begin exact ⟨λ k b hb, k _ _ hb rfl, λ k a' b hb h, h ▸ k _ hb⟩, end +alias finset.sup_le_iff ↔ _ sup_le + +attribute [protected] sup_le + +lemma le_sup {b : β} (hb : b ∈ s) : f b ≤ s.sup f := finset.sup_le_iff.1 le_rfl _ hb + @[simp] lemma sup_bUnion [decidable_eq β] (s : finset γ) (t : γ → finset β) : (s.bUnion t).sup f = s.sup (λ x, (t x).sup f) := eq_of_forall_ge_iff $ λ c, by simp [@forall_swap _ β] @@ -91,19 +97,12 @@ lemma sup_ite (p : β → Prop) [decidable_pred p] : (s.filter p).sup f ⊔ (s.filter (λ i, ¬ p i)).sup g := fold_ite _ -lemma sup_le {a : α} : (∀ b ∈ s, f b ≤ a) → s.sup f ≤ a := -finset.sup_le_iff.2 - -lemma le_sup {b : β} (hb : b ∈ s) : f b ≤ s.sup f := -finset.sup_le_iff.1 le_rfl _ hb - lemma sup_mono_fun {g : β → α} (h : ∀ b ∈ s, f b ≤ g b) : s.sup f ≤ s.sup g := -sup_le (λ b hb, le_trans (h b hb) (le_sup hb)) +finset.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) +lemma sup_mono (h : s₁ ⊆ s₂) : s₁.sup f ≤ s₂.sup f := finset.sup_le $ λ b hb, le_sup $ h hb -lemma sup_comm (s : finset β) (t : finset γ) (f : β → γ → α) : +protected 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 refine eq_of_forall_ge_iff (λ a, _), @@ -118,17 +117,13 @@ end lemma sup_product_left (s : finset β) (t : finset γ) (f : β × γ → α) : (s ×ˢ t).sup f = s.sup (λ i, t.sup $ λ i', f ⟨i, i'⟩) := begin - refine le_antisymm _ (sup_le (λ i hi, sup_le $ λ i' hi', le_sup $ mem_product.2 ⟨hi, hi'⟩)), - refine sup_le _, - rintro ⟨i, i'⟩ hi, - rw mem_product at hi, - refine le_trans _ (le_sup hi.1), - convert le_sup hi.2, + simp only [le_antisymm_iff, finset.sup_le_iff, mem_product, and_imp, prod.forall], + exact ⟨λ b c hb hc, (le_sup hb).trans' $ le_sup hc, λ b hb c hc, le_sup $ mem_product.2 ⟨hb, hc⟩⟩, end lemma sup_product_right (s : finset β) (t : finset γ) (f : β × γ → α) : (s ×ˢ t).sup f = t.sup (λ i', s.sup $ λ i, f ⟨i, i'⟩) := -by rw [sup_product_left, sup_comm] +by rw [sup_product_left, finset.sup_comm] @[simp] lemma sup_erase_bot [decidable_eq α] (s : finset α) : (s.erase ⊥).sup id = s.sup id := begin @@ -299,27 +294,26 @@ lemma inf_const {s : finset β} (h : s.nonempty) (c : α) : s.inf (λ _, c) = c @[simp] lemma inf_top (s : finset β) : s.inf (λ _, ⊤) = (⊤ : α) := @sup_bot αᵒᵈ _ _ _ _ -lemma le_inf_iff {a : α} : a ≤ s.inf f ↔ ∀ b ∈ s, a ≤ f b := +protected lemma le_inf_iff {a : α} : a ≤ s.inf f ↔ ∀ b ∈ s, a ≤ f b := @finset.sup_le_iff αᵒᵈ _ _ _ _ _ _ -lemma inf_le {b : β} (hb : b ∈ s) : s.inf f ≤ f b := -le_inf_iff.1 le_rfl _ hb +alias finset.le_inf_iff ↔ _ le_inf + +attribute [protected] le_inf -lemma le_inf {a : α} : (∀ b ∈ s, a ≤ f b) → a ≤ s.inf f := -le_inf_iff.2 +lemma inf_le {b : β} (hb : b ∈ s) : s.inf f ≤ f b := finset.le_inf_iff.1 le_rfl _ hb lemma inf_mono_fun {g : β → α} (h : ∀ b ∈ s, f b ≤ g b) : s.inf f ≤ s.inf g := -le_inf (λ b hb, le_trans (inf_le hb) (h b hb)) +finset.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) +lemma inf_mono (h : s₁ ⊆ s₂) : s₂.inf f ≤ s₁.inf f := finset.le_inf $ λ b hb, inf_le $h hb lemma inf_attach (s : finset β) (f : β → α) : s.attach.inf (λ x, f x) = s.inf f := @sup_attach αᵒᵈ _ _ _ _ _ -lemma inf_comm (s : finset β) (t : finset γ) (f : β → γ → α) : +protected 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 αᵒᵈ _ _ _ _ _ _ _ +@finset.sup_comm αᵒᵈ _ _ _ _ _ _ _ lemma inf_product_left (s : finset β) (t : finset γ) (f : β × γ → α) : (s ×ˢ t).inf f = s.inf (λ i, t.inf $ λ i', f ⟨i, i'⟩) := @@ -535,7 +529,8 @@ by { rw ←with_bot.coe_eq_coe, simp only [coe_sup', sup_insert, with_bot.coe_su ({b} : finset β).sup' h f = f b := rfl lemma sup'_le {a : α} (hs : ∀ b ∈ s, f b ≤ a) : s.sup' H f ≤ a := -by { rw [←with_bot.coe_le_coe, coe_sup'], exact sup_le (λ b h, with_bot.coe_le_coe.2 $ hs b h), } +by { rw [←with_bot.coe_le_coe, coe_sup'], + exact finset.sup_le (λ b h, with_bot.coe_le_coe.2 $ hs b h) } lemma le_sup' {b : β} (h : b ∈ s) : f b ≤ s.sup' ⟨b, h⟩ f := by { rw [←with_bot.coe_le_coe, coe_sup'], exact le_sup h, } @@ -673,7 +668,7 @@ section sup variables [semilattice_sup α] [order_bot α] lemma sup'_eq_sup {s : finset β} (H : s.nonempty) (f : β → α) : s.sup' H f = s.sup f := -le_antisymm (sup'_le H f (λ b, le_sup)) (sup_le (λ b, le_sup' f)) +le_antisymm (sup'_le H f (λ b, le_sup)) (finset.sup_le (λ b, le_sup' f)) lemma sup_closed_of_sup_closed {s : set α} (t : finset α) (htne : t.nonempty) (h_subset : ↑t ⊆ s) (h : ∀ a b ∈ s, a ⊔ b ∈ s) : t.sup id ∈ s := @@ -860,9 +855,9 @@ finset.not_mem_of_max_lt_coe $ h₂.trans_lt $ with_bot.coe_lt_coe.mpr h₁ lemma max_mono {s t : finset α} (st : s ⊆ t) : s.max ≤ t.max := sup_mono st -lemma max_le {M : with_bot α} {s : finset α} (st : ∀ a : α, a ∈ s → (a : with_bot α) ≤ M) : +protected lemma max_le {M : with_bot α} {s : finset α} (st : ∀ a ∈ s, (a : with_bot α) ≤ M) : s.max ≤ M := -sup_le st +finset.sup_le st /-- Let `s` be a finset in a linear order. Then `s.min` is the minimum of `s` if `s` is not empty, and `⊤` otherwise. It belongs to `with_top α`. If you want to get an element of `α`, see @@ -910,9 +905,9 @@ finset.not_mem_of_coe_lt_min $ (with_top.coe_lt_coe.mpr h₁).trans_eq h₂.symm lemma min_mono {s t : finset α} (st : s ⊆ t) : t.min ≤ s.min := inf_mono st -lemma le_min {m : with_top α} {s : finset α} (st : ∀ a : α, a ∈ s → m ≤ a) : +protected lemma le_min {m : with_top α} {s : finset α} (st : ∀ a : α, a ∈ s → m ≤ a) : m ≤ s.min := -le_inf st +finset.le_inf st /-- Given a nonempty finset `s` in a linear order `α`, then `s.min' h` is its minimum, as an element of `α`, where `h` is a proof of nonemptiness. Without this assumption, use instead `s.min`, diff --git a/src/data/finset/sigma.lean b/src/data/finset/sigma.lean index 6be4714fa13d1..6c463ce72a0fb 100644 --- a/src/data/finset/sigma.lean +++ b/src/data/finset/sigma.lean @@ -77,11 +77,8 @@ variables (s t) (f : (Σ i, α i) → β) lemma sup_sigma [semilattice_sup β] [order_bot β] : (s.sigma t).sup f = s.sup (λ i, (t i).sup $ λ b, f ⟨i, b⟩) := begin - refine (sup_le _).antisymm (sup_le $ λ i hi, sup_le $ λ b hb, le_sup $ mem_sigma.2 ⟨hi, hb⟩), - rintro ⟨i, b⟩ hb, - rw mem_sigma at hb, - refine le_trans _ (le_sup hb.1), - convert le_sup hb.2, + simp only [le_antisymm_iff, finset.sup_le_iff, mem_sigma, and_imp, sigma.forall], + exact ⟨λ i a hi ha, (le_sup hi).trans' $ le_sup ha, λ i hi a ha, le_sup $ mem_sigma.2 ⟨hi, ha⟩⟩, end lemma inf_sigma [semilattice_inf β] [order_top β] : diff --git a/src/data/polynomial/degree/lemmas.lean b/src/data/polynomial/degree/lemmas.lean index ed58c06990a27..f9f10db9a9563 100644 --- a/src/data/polynomial/degree/lemmas.lean +++ b/src/data/polynomial/degree/lemmas.lean @@ -33,7 +33,7 @@ else with_bot.coe_le_coe.1 $ calc ↑(nat_degree (p.comp q)) = degree (p.comp q) : (degree_eq_nat_degree h0).symm ... = _ : congr_arg degree comp_eq_sum_left ... ≤ _ : degree_sum_le _ _ - ... ≤ _ : sup_le (λ n hn, + ... ≤ _ : finset.sup_le (λ n hn, calc degree (C (coeff p n) * q ^ n) ≤ degree (C (coeff p n)) + degree (q ^ n) : degree_mul_le _ _ ... ≤ nat_degree (C (coeff p n)) + n • (degree q) : diff --git a/src/data/polynomial/degree/trailing_degree.lean b/src/data/polynomial/degree/trailing_degree.lean index 3974f895e9ba4..3903abea51ee5 100644 --- a/src/data/polynomial/degree/trailing_degree.lean +++ b/src/data/polynomial/degree/trailing_degree.lean @@ -273,7 +273,7 @@ end lemma le_trailing_degree_mul : p.trailing_degree + q.trailing_degree ≤ (p * q).trailing_degree := begin - refine le_min (λ n hn, _), + refine finset.le_min (λ n hn, _), rw [mem_support_iff, coeff_mul] at hn, obtain ⟨⟨i, j⟩, hij, hpq⟩ := exists_ne_zero_of_sum_ne_zero hn, refine (add_le_add (min_le (mem_support_iff.mpr (left_ne_zero_of_mul hpq))) diff --git a/src/data/polynomial/derivative.lean b/src/data/polynomial/derivative.lean index d159946d9c3c8..cc33b32a29fed 100644 --- a/src/data/polynomial/derivative.lean +++ b/src/data/polynomial/derivative.lean @@ -292,7 +292,7 @@ begin simp [hp] }, apply le_antisymm, { rw derivative_apply, - apply le_trans (degree_sum_le _ _) (sup_le (λ n hn, _)), + apply le_trans (degree_sum_le _ _) (finset.sup_le (λ n hn, _)), apply le_trans (degree_C_mul_X_pow_le _ _) (with_bot.coe_le_coe.2 (tsub_le_tsub_right _ _)), apply le_nat_degree_of_mem_supp _ hn }, { refine le_sup _, diff --git a/src/data/polynomial/unit_trinomial.lean b/src/data/polynomial/unit_trinomial.lean index 0a9f46f7d7edd..08a659408f6c1 100644 --- a/src/data/polynomial/unit_trinomial.lean +++ b/src/data/polynomial/unit_trinomial.lean @@ -56,8 +56,8 @@ by rw [trinomial_def, coeff_add, coeff_add, coeff_C_mul_X_pow, coeff_C_mul_X_pow lemma trinomial_nat_degree (hkm : k < m) (hmn : m < n) (hw : w ≠ 0) : (trinomial k m n u v w).nat_degree = n := begin - refine nat_degree_eq_of_degree_eq_some (le_antisymm (sup_le (λ i h, _)) - (le_degree_of_ne_zero (by rwa trinomial_leading_coeff' hkm hmn))), + refine nat_degree_eq_of_degree_eq_some ((finset.sup_le $ λ i h, _).antisymm $ + le_degree_of_ne_zero $ by rwa trinomial_leading_coeff' hkm hmn), replace h := support_trinomial' k m n u v w h, rw [mem_insert, mem_insert, mem_singleton] at h, rcases h with rfl | rfl | rfl, @@ -69,8 +69,8 @@ end lemma trinomial_nat_trailing_degree (hkm : k < m) (hmn : m < n) (hu : u ≠ 0) : (trinomial k m n u v w).nat_trailing_degree = k := begin - refine nat_trailing_degree_eq_of_trailing_degree_eq_some (le_antisymm (le_inf (λ i h, _)) - (le_trailing_degree_of_ne_zero (by rwa trinomial_trailing_coeff' hkm hmn))).symm, + refine nat_trailing_degree_eq_of_trailing_degree_eq_some ((finset.le_inf $ λ i h, _).antisymm $ + le_trailing_degree_of_ne_zero $ by rwa trinomial_trailing_coeff' hkm hmn).symm, replace h := support_trinomial' k m n u v w h, rw [mem_insert, mem_insert, mem_singleton] at h, rcases h with rfl | rfl | rfl, From d012cd09a9b256d870751284dd6a29882b0be105 Mon Sep 17 00:00:00 2001 From: sgouezel Date: Thu, 15 Dec 2022 18:06:20 +0000 Subject: [PATCH 0060/1291] feat(measure_theory/covering) weaken sigma_compact_space to second_countable_topology (#17960) Currently, theorems on differentiation of measures require the assumption that the space is sigma compact. However, they all hold under the weaker assumption that the space is second-countable, as shown by this PR. --- src/measure_theory/covering/besicovitch.lean | 11 ++--- .../covering/density_theorem.lean | 4 +- .../covering/differentiation.lean | 25 +++++++--- .../covering/liminf_limsup.lean | 5 +- src/measure_theory/measure/measure_space.lean | 48 +++++++++++++++++++ src/measure_theory/measure/regular.lean | 22 ++++----- 6 files changed, 84 insertions(+), 31 deletions(-) diff --git a/src/measure_theory/covering/besicovitch.lean b/src/measure_theory/covering/besicovitch.lean index f0576fecbea21..186470555d8a2 100644 --- a/src/measure_theory/covering/besicovitch.lean +++ b/src/measure_theory/covering/besicovitch.lean @@ -1109,7 +1109,7 @@ begin { exact closed_ball_subset_closed_ball hr.2 } end -variables [metric_space β] [measurable_space β] [borel_space β] [sigma_compact_space β] +variables [metric_space β] [measurable_space β] [borel_space β] [second_countable_topology β] [has_besicovitch_covering β] /-- In a space with the Besicovitch covering property, the ratio of the measure of balls converges @@ -1119,7 +1119,6 @@ lemma ae_tendsto_rn_deriv ∀ᵐ x ∂μ, tendsto (λ r, ρ (closed_ball x r) / μ (closed_ball x r)) (𝓝[>] 0) (𝓝 (ρ.rn_deriv μ x)) := begin - haveI : second_countable_topology β := emetric.second_countable_of_sigma_compact β, filter_upwards [vitali_family.ae_tendsto_rn_deriv (besicovitch.vitali_family μ) ρ] with x hx, exact hx.comp (tendsto_filter_at μ x) end @@ -1134,7 +1133,6 @@ lemma ae_tendsto_measure_inter_div_of_measurable_set ∀ᵐ x ∂μ, tendsto (λ r, μ (s ∩ closed_ball x r) / μ (closed_ball x r)) (𝓝[>] 0) (𝓝 (s.indicator 1 x)) := begin - haveI : second_countable_topology β := emetric.second_countable_of_sigma_compact β, filter_upwards [vitali_family.ae_tendsto_measure_inter_div_of_measurable_set (besicovitch.vitali_family μ) hs], assume x hx, @@ -1150,10 +1148,7 @@ See also `is_doubling_measure.ae_tendsto_measure_inter_div`. -/ lemma ae_tendsto_measure_inter_div (μ : measure β) [is_locally_finite_measure μ] (s : set β) : ∀ᵐ x ∂(μ.restrict s), tendsto (λ r, μ (s ∩ (closed_ball x r)) / μ (closed_ball x r)) (𝓝[>] 0) (𝓝 1) := -begin - haveI : second_countable_topology β := emetric.second_countable_of_sigma_compact β, - filter_upwards [vitali_family.ae_tendsto_measure_inter_div (besicovitch.vitali_family μ)] - with x hx using hx.comp (tendsto_filter_at μ x), -end +by filter_upwards [vitali_family.ae_tendsto_measure_inter_div (besicovitch.vitali_family μ)] + with x hx using hx.comp (tendsto_filter_at μ x) end besicovitch diff --git a/src/measure_theory/covering/density_theorem.lean b/src/measure_theory/covering/density_theorem.lean index 4e5cc91bedc99..a17f343fa8511 100644 --- a/src/measure_theory/covering/density_theorem.lean +++ b/src/measure_theory/covering/density_theorem.lean @@ -29,8 +29,6 @@ noncomputable theory open set filter metric measure_theory topological_space open_locale nnreal topological_space -local attribute [instance] emetric.second_countable_of_sigma_compact - namespace is_doubling_measure variables {α : Type*} [metric_space α] [measurable_space α] (μ : measure α) [is_doubling_measure μ] @@ -135,7 +133,7 @@ end end section applications -variables [sigma_compact_space α] [borel_space α] [is_locally_finite_measure μ] +variables [second_countable_topology α] [borel_space α] [is_locally_finite_measure μ] {E : Type*} [normed_add_comm_group E] /-- A version of *Lebesgue's density theorem* for a sequence of closed balls whose centers are diff --git a/src/measure_theory/covering/differentiation.lean b/src/measure_theory/covering/differentiation.lean index 80b14a129b52c..1f9abd2f96386 100644 --- a/src/measure_theory/covering/differentiation.lean +++ b/src/measure_theory/covering/differentiation.lean @@ -13,8 +13,8 @@ import measure_theory.decomposition.lebesgue /-! # Differentiation of measures -On a metric space with a measure `μ`, consider a Vitali family (i.e., for each `x` one has a family -of sets shrinking to `x`, with a good behavior with respect to covering theorems). +On a second countable metric space with a measure `μ`, consider a Vitali family (i.e., for each `x` +one has a family of sets shrinking to `x`, with a good behavior with respect to covering theorems). Consider also another measure `ρ`. Then, for almost every `x`, the ratio `ρ a / μ a` converges when `a` shrinks to `x` along the Vitali family, towards the Radon-Nikodym derivative of `ρ` with respect to `μ`. This is the main theorem on differentiation of measures. @@ -56,6 +56,19 @@ almost everywhere measurable, again based on the disjoint subcovering argument (see `vitali_family.exists_measurable_supersets_lim_ratio`), and then proceed as sketched above but replacing `v.lim_ratio ρ` by a measurable version called `v.lim_ratio_meas ρ`. +## Counterexample + +The standing assumption in this file is that spaces are second countable. Without this assumption, +measures may be zero locally but nonzero globally, which is not compatible with differentiation +theory (which deduces global information from local one). Here is an example displaying this +behavior. + +Define a measure `μ` by `μ s = 0` if `s` is covered by countably many balls of radius `1`, +and `μ s = ∞` otherwise. This is indeed a countably additive measure, which is moreover +locally finite and doubling at small scales. It vanishes on every ball of radius `1`, so all the +quantities in differentiation theory (defined as ratios of measures as the radius tends to zero) +make no sense. However, the measure is not globally zero if the space is big enough. + ## References * [Herbert Federer, Geometric Measure Theory, Chapter 2.9][Federer1996] @@ -64,8 +77,6 @@ but replacing `v.lim_ratio ρ` by a measurable version called `v.lim_ratio_meas open measure_theory metric set filter topological_space measure_theory.measure open_locale filter ennreal measure_theory nnreal topological_space -local attribute [instance] emetric.second_countable_of_sigma_compact - variables {α : Type*} [metric_space α] {m0 : measurable_space α} {μ : measure α} (v : vitali_family μ) {E : Type*} [normed_add_comm_group E] @@ -113,7 +124,7 @@ end /-- If two measures `ρ` and `ν` have, at every point of a set `s`, arbitrarily small sets in a Vitali family satisfying `ρ a ≤ ν a`, then `ρ s ≤ ν s` if `ρ ≪ μ`.-/ -theorem measure_le_of_frequently_le [sigma_compact_space α] [borel_space α] +theorem measure_le_of_frequently_le [second_countable_topology α] [borel_space α] {ρ : measure α} (ν : measure α) [is_locally_finite_measure ν] (hρ : ρ ≪ μ) (s : set α) (hs : ∀ x ∈ s, ∃ᶠ a in v.filter_at x, ρ a ≤ ν a) : ρ s ≤ ν s := @@ -141,7 +152,7 @@ end section -variables [sigma_compact_space α] [borel_space α] [is_locally_finite_measure μ] +variables [second_countable_topology α] [borel_space α] [is_locally_finite_measure μ] {ρ : measure α} [is_locally_finite_measure ρ] /-- If a measure `ρ` is singular with respect to `μ`, then for `μ` almost every `x`, the ratio @@ -797,7 +808,7 @@ begin A minor technical inconvenience is that constants are not integrable, so to apply previous lemmas we need to replace `c` with the restriction of `c` to a finite measure set `A n` in the above sketch. -/ - let A := measure_theory.measure.finite_spanning_sets_in_open μ, + let A := measure_theory.measure.finite_spanning_sets_in_open' μ, rcases h'f.is_separable_range with ⟨t, t_count, ht⟩, have main : ∀ᵐ x ∂μ, ∀ (n : ℕ) (c : E) (hc : c ∈ t), tendsto (λ a, (∫⁻ y in a, ‖f y - (A.set n).indicator (λ y, c) y‖₊ ∂μ) / μ a) diff --git a/src/measure_theory/covering/liminf_limsup.lean b/src/measure_theory/covering/liminf_limsup.lean index 498f872440ae8..11eb21e62c41e 100644 --- a/src/measure_theory/covering/liminf_limsup.lean +++ b/src/measure_theory/covering/liminf_limsup.lean @@ -21,10 +21,11 @@ carrying a doubling measure. -/ -open set filter metric measure_theory +open set filter metric measure_theory topological_space open_locale nnreal ennreal topological_space -variables {α : Type*} [metric_space α] [sigma_compact_space α] [measurable_space α] [borel_space α] +variables {α : Type*} [metric_space α] [second_countable_topology α] [measurable_space α] + [borel_space α] variables (μ : measure α) [is_locally_finite_measure μ] [is_doubling_measure μ] /-- This is really an auxiliary result en route to `blimsup_cthickening_ae_le_of_eventually_mul_le` diff --git a/src/measure_theory/measure/measure_space.lean b/src/measure_theory/measure/measure_space.lean index 68716ac6dd571..8c561cdb00abd 100644 --- a/src/measure_theory/measure/measure_space.lean +++ b/src/measure_theory/measure/measure_space.lean @@ -3366,6 +3366,17 @@ begin ennreal.coe_of_nnreal_hom, ne.def, not_false_iff], end +protected lemma measure.is_topological_basis_is_open_lt_top [topological_space α] (μ : measure α) + [is_locally_finite_measure μ] : + topological_space.is_topological_basis {s | is_open s ∧ μ s < ∞} := +begin + refine topological_space.is_topological_basis_of_open_of_nhds (λ s hs, hs.1) _, + assume x s xs hs, + rcases μ.exists_is_open_measure_lt_top x with ⟨v, xv, hv, μv⟩, + refine ⟨v ∩ s, ⟨hv.inter hs, lt_of_le_of_lt _ μv⟩, ⟨xv, xs⟩, inter_subset_right _ _⟩, + exact measure_mono (inter_subset_left _ _), +end + /-- A measure `μ` is finite on compacts if any compact set `K` satisfies `μ K < ∞`. -/ @[protect_proj] class is_finite_measure_on_compacts [topological_space α] (μ : measure α) : Prop := (lt_top_of_is_compact : ∀ ⦃K : set α⦄, is_compact K → μ K < ∞) @@ -3918,6 +3929,43 @@ def measure_theory.measure.finite_spanning_sets_in_open [topological_space α] ((is_compact_compact_covering α n).exists_open_superset_measure_lt_top μ).some_spec.fst) (Union_compact_covering α) } +open topological_space + +/-- A locally finite measure on a second countable topological space admits a finite spanning +sequence of open sets. -/ +@[irreducible] def measure_theory.measure.finite_spanning_sets_in_open' [topological_space α] + [second_countable_topology α] {m : measurable_space α} (μ : measure α) + [is_locally_finite_measure μ] : + μ.finite_spanning_sets_in {K | is_open K} := +begin + suffices H : nonempty (μ.finite_spanning_sets_in {K | is_open K}), from H.some, + casesI is_empty_or_nonempty α, + { exact + ⟨{ set := λ n, ∅, set_mem := λ n, by simp, finite := λ n, by simp, spanning := by simp }⟩ }, + inhabit α, + let S : set (set α) := {s | is_open s ∧ μ s < ∞}, + obtain ⟨T, T_count, TS, hT⟩ : ∃ T : set (set α), T.countable ∧ T ⊆ S ∧ ⋃₀ T = ⋃₀ S := + is_open_sUnion_countable S (λ s hs, hs.1), + rw μ.is_topological_basis_is_open_lt_top.sUnion_eq at hT, + have T_ne : T.nonempty, + { by_contra h'T, + simp only [not_nonempty_iff_eq_empty.1 h'T, sUnion_empty] at hT, + simpa only [← hT] using mem_univ (default : α) }, + obtain ⟨f, hf⟩ : ∃ f : ℕ → set α, T = range f, from T_count.exists_eq_range T_ne, + have fS : ∀ n, f n ∈ S, + { assume n, + apply TS, + rw hf, + exact mem_range_self n }, + refine ⟨{ set := f, set_mem := λ n, (fS n).1, finite := λ n, (fS n).2, spanning := _ }⟩, + apply eq_univ_of_forall (λ x, _), + obtain ⟨t, tT, xt⟩ : ∃ (t : set α), t ∈ range f ∧ x ∈ t, + { have : x ∈ ⋃₀ T, by simp only [hT], + simpa only [mem_sUnion, exists_prop, ← hf] }, + obtain ⟨n, rfl⟩ : ∃ (n : ℕ), f n = t, by simpa only using tT, + exact mem_Union_of_mem _ xt, +end + section measure_Ixx variables [preorder α] [topological_space α] [compact_Icc_space α] diff --git a/src/measure_theory/measure/regular.lean b/src/measure_theory/measure/regular.lean index 3b2ffe5999168..520260f52a41a 100644 --- a/src/measure_theory/measure/regular.lean +++ b/src/measure_theory/measure/regular.lean @@ -66,12 +66,10 @@ is automatically satisfied by any finite measure on a metric space. * `set.measure_eq_infi_is_open` asserts that, when `μ` is outer regular, the measure of a set is the infimum of the measure of open sets containing it. -* `set.exists_is_open_lt_of_lt'` asserts that, when `μ` is outer regular, for every set `s` +* `set.exists_is_open_lt_of_lt` asserts that, when `μ` is outer regular, for every set `s` and `r > μ s` there exists an open superset `U ⊇ s` of measure less than `r`. * push forward of an outer regular measure is outer regular, and scalar multiplication of a regular measure by a finite number is outer regular. -* `measure_theory.measure.outer_regular.of_sigma_compact_space_of_is_locally_finite_measure`: - a locally finite measure on a `σ`-compact metric (or even pseudo emetric) space is outer regular. ### Weakly regular measures @@ -87,9 +85,9 @@ is automatically satisfied by any finite measure on a metric space. * `measure_theory.measure.weakly_regular.of_pseudo_emetric_space_of_is_finite_measure` is an instance registering that a finite measure on a metric space is weakly regular (in fact, a pseudo emetric space is enough); -* `measure_theory.measure.weakly_regular.of_pseudo_emetric_sigma_compact_space_of_locally_finite` - is an instance registering that a locally finite measure on a `σ`-compact metric space (or even - a pseudo emetric space) is weakly regular. +* `measure_theory.measure.weakly_regular.of_pseudo_emetric_second_countable_of_locally_finite` + is an instance registering that a locally finite measure on a second countable metric space (or + even a pseudo emetric space) is weakly regular. ### Regular measures @@ -613,16 +611,16 @@ instance of_pseudo_emetric_space_of_is_finite_measure {X : Type*} [pseudo_emetri weakly_regular μ := (inner_regular.of_pseudo_emetric_space μ).weakly_regular_of_finite μ -/-- Any locally finite measure on a `σ`-compact metric space (or even a pseudo emetric space) is -weakly regular. -/ +/-- Any locally finite measure on a second countable metric space (or even a pseudo emetric space) +is weakly regular. -/ @[priority 100] -- see Note [lower instance priority] -instance of_pseudo_emetric_sigma_compact_space_of_locally_finite {X : Type*} - [pseudo_emetric_space X] [sigma_compact_space X] [measurable_space X] [borel_space X] +instance of_pseudo_emetric_second_countable_of_locally_finite {X : Type*} [pseudo_emetric_space X] + [topological_space.second_countable_topology X] [measurable_space X] [borel_space X] (μ : measure X) [is_locally_finite_measure μ] : weakly_regular μ := begin haveI : outer_regular μ, - { refine (μ.finite_spanning_sets_in_open.mono' $ λ U hU, _).outer_regular, + { refine (μ.finite_spanning_sets_in_open'.mono' $ λ U hU, _).outer_regular, haveI : fact (μ U < ∞), from ⟨hU.2⟩, exact ⟨hU.1, infer_instance⟩ }, exact ⟨inner_regular.of_pseudo_emetric_space μ⟩ @@ -630,6 +628,8 @@ end end weakly_regular +local attribute [instance] emetric.second_countable_of_sigma_compact + /-- Any locally finite measure on a `σ`-compact (e)metric space is regular. -/ @[priority 100] -- see Note [lower instance priority] instance regular.of_sigma_compact_space_of_is_locally_finite_measure {X : Type*} From b3f25363ae62cb169e72cd6b8b1ac97bacf21ca7 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 16 Dec 2022 00:10:50 +0000 Subject: [PATCH 0061/1291] chore(order/monovary): Move (#17946) Now that we have an `order.monotone` subfolder, `monovary` belongs there. --- src/algebra/order/rearrangement.lean | 2 +- src/order/{ => monotone}/monovary.lean | 0 2 files changed, 1 insertion(+), 1 deletion(-) rename src/order/{ => monotone}/monovary.lean (100%) diff --git a/src/algebra/order/rearrangement.lean b/src/algebra/order/rearrangement.lean index f3fa11c25eb67..f4575a8d4414c 100644 --- a/src/algebra/order/rearrangement.lean +++ b/src/algebra/order/rearrangement.lean @@ -7,7 +7,7 @@ import algebra.big_operators.basic import algebra.order.module import data.prod.lex import group_theory.perm.support -import order.monovary +import order.monotone.monovary import tactic.abel /-! diff --git a/src/order/monovary.lean b/src/order/monotone/monovary.lean similarity index 100% rename from src/order/monovary.lean rename to src/order/monotone/monovary.lean From d5ca915d3b1c72307a7b151fe6ee72c886813a12 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 16 Dec 2022 11:04:42 +0000 Subject: [PATCH 0062/1291] refactor(measure_theory/group/fundamental_domain): Use `pairwise` (#17928) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Replace `∀ g ≠ (1 : G), ae_disjoint μ (g • s) s` ("translates are disjoint to the original set") by the stronger and more symmetric `pairwise (ae_disjoint μ on λ g : G, g • s)` ("translates are disjoint"). In full generality, the latter condition is the one we mean, and indeed this change reduces typeclass assumptions for a few lemmas. --- .../group/fundamental_domain.lean | 95 ++++++++++--------- 1 file changed, 49 insertions(+), 46 deletions(-) diff --git a/src/measure_theory/group/fundamental_domain.lean b/src/measure_theory/group/fundamental_domain.lean index a22cb7b53bd5d..75d690a89c0ba 100644 --- a/src/measure_theory/group/fundamental_domain.lean +++ b/src/measure_theory/group/fundamental_domain.lean @@ -38,7 +38,7 @@ a.e. disjoint and cover the whole space. -/ [has_vadd G α] [measurable_space α] (s : set α) (μ : measure α . volume_tac) : Prop := (null_measurable_set : null_measurable_set s μ) (ae_covers : ∀ᵐ x ∂μ, ∃ g : G, g +ᵥ x ∈ s) -(ae_disjoint : ∀ g ≠ (0 : G), ae_disjoint μ (g +ᵥ s) s) +(ae_disjoint : pairwise $ ae_disjoint μ on λ g : G, g +ᵥ s) /-- A measurable set `s` is a *fundamental domain* for an action of a group `G` on a measurable space `α` with respect to a measure `α` if the sets `g • s`, `g : G`, are pairwise a.e. disjoint and @@ -48,12 +48,13 @@ structure is_fundamental_domain (G : Type*) {α : Type*} [has_one G] [has_smul G [measurable_space α] (s : set α) (μ : measure α . volume_tac) : Prop := (null_measurable_set : null_measurable_set s μ) (ae_covers : ∀ᵐ x ∂μ, ∃ g : G, g • x ∈ s) -(ae_disjoint : ∀ g ≠ (1 : G), ae_disjoint μ (g • s) s) +(ae_disjoint : pairwise $ ae_disjoint μ on λ g : G, g • s) -namespace is_fundamental_domain +variables {G H α β E : Type*} -variables {G H α β E : Type*} [group G] [group H] [mul_action G α] [measurable_space α] - [mul_action H β] [measurable_space β] [normed_add_comm_group E] {s t : set α} {μ : measure α} +namespace is_fundamental_domain +variables [group G] [group H] [mul_action G α] [measurable_space α] [mul_action H β] + [measurable_space β] [normed_add_comm_group E] {s t : set α} {μ : measure α} /-- If for each `x : α`, exactly one of `g • x`, `g : G`, belongs to a measurable set `s`, then `s` is a fundamental domain for the action of `G` on `α`. -/ @@ -63,13 +64,23 @@ lemma mk' (h_meas : null_measurable_set s μ) (h_exists : ∀ x : α, ∃! g : G is_fundamental_domain G s μ := { null_measurable_set := h_meas, ae_covers := eventually_of_forall $ λ x, (h_exists x).exists, - ae_disjoint := λ g hne, disjoint.ae_disjoint $ disjoint_left.2 + ae_disjoint := λ a b hab, disjoint.ae_disjoint $ disjoint_left.2 $ λ x hxa hxb, begin - rintro _ ⟨x, hx, rfl⟩ hgx, - rw ← one_smul G x at hx, - exact hne ((h_exists x).unique hgx hx) + rw mem_smul_set_iff_inv_smul_mem at hxa hxb, + exact hab (inv_injective $ (h_exists x).unique hxa hxb), end } +/-- For `s` to be a fundamental domain, it's enough to check `ae_disjoint (g • s) s` for `g ≠ 1`. -/ +@[to_additive "For `s` to be a fundamental domain, it's enough to check `ae_disjoint (g +ᵥ s) s` for +`g ≠ 0`."] +lemma mk'' (h_meas : null_measurable_set s μ) (h_ae_covers : ∀ᵐ x ∂μ, ∃ g : G, g • x ∈ s) + (h_ae_disjoint : ∀ g ≠ (1 : G), ae_disjoint μ (g • s) s) + (h_qmp : ∀ (g : G), quasi_measure_preserving ((•) g : α → α) μ μ) : + is_fundamental_domain G s μ := +{ null_measurable_set := h_meas, + ae_covers := h_ae_covers, + ae_disjoint := pairwise_ae_disjoint_of_ae_disjoint_forall_ne_one h_ae_disjoint h_qmp } + /-- If a measurable space has a finite measure `μ` and a countable group `G` acts quasi-measure-preservingly, then to show that a set `s` is a fundamental domain, it is sufficient to check that its translates `g • s` are (almost) disjoint and that the sum `∑' g, μ (g • s)` is @@ -85,12 +96,12 @@ lemma mk_of_measure_univ_le [is_finite_measure μ] [countable G] (h_qmp : ∀ (g : G), quasi_measure_preserving ((•) g : α → α) μ μ) (h_measure_univ_le : μ (univ : set α) ≤ ∑' (g : G), μ (g • s)) : is_fundamental_domain G s μ := +have ae_disjoint : pairwise (ae_disjoint μ on (λ (g : G), g • s)) := + pairwise_ae_disjoint_of_ae_disjoint_forall_ne_one h_ae_disjoint h_qmp, { null_measurable_set := h_meas, - ae_disjoint := h_ae_disjoint, + ae_disjoint := ae_disjoint, ae_covers := begin - replace ae_disjoint : pairwise (ae_disjoint μ on (λ (g : G), g • s)) := - pairwise_ae_disjoint_of_ae_disjoint_forall_ne_one h_ae_disjoint h_qmp, replace h_meas : ∀ (g : G), null_measurable_set (g • s) μ := λ g, by { rw [← inv_inv g, ← preimage_smul], exact h_meas.preimage (h_qmp g⁻¹), }, have h_meas' : null_measurable_set {a | ∃ (g : G), g • a ∈ s} μ, @@ -107,26 +118,7 @@ eventually_eq_univ.2 $ h.ae_covers.mono $ λ x ⟨g, hg⟩, mem_Union.2 ⟨g⁻ @[to_additive] lemma mono (h : is_fundamental_domain G s μ) {ν : measure α} (hle : ν ≪ μ) : is_fundamental_domain G s ν := -⟨h.1.mono_ac hle, hle h.2, λ g hg, hle (h.3 g hg)⟩ - -variables [measurable_space G] [has_measurable_smul G α] [smul_invariant_measure G α μ] - -@[to_additive] lemma null_measurable_set_smul (h : is_fundamental_domain G s μ) (g : G) : - null_measurable_set (g • s) μ := -h.null_measurable_set.smul g - -@[to_additive] lemma restrict_restrict (h : is_fundamental_domain G s μ) (g : G) (t : set α) : - (μ.restrict t).restrict (g • s) = μ.restrict (g • s ∩ t) := -restrict_restrict₀ ((h.null_measurable_set_smul g).mono restrict_le_self) - -@[to_additive] lemma pairwise_ae_disjoint (h : is_fundamental_domain G s μ) : - pairwise (λ g₁ g₂ : G, ae_disjoint μ (g₁ • s) (g₂ • s)) := -pairwise_ae_disjoint_of_ae_disjoint_forall_ne_one h.ae_disjoint - (λ g, measure_preserving.quasi_measure_preserving $ by simp) - -@[to_additive] lemma pairwise_ae_disjoint_of_ac {ν} (h : is_fundamental_domain G s μ) (hν : ν ≪ μ) : - pairwise (λ g₁ g₂ : G, ae_disjoint ν (g₁ • s) (g₂ • s)) := -h.pairwise_ae_disjoint.mono $ λ g₁ g₂ H, hν H +⟨h.1.mono_ac hle, hle h.2, h.ae_disjoint.mono $ λ a b hab, hle hab⟩ @[to_additive] lemma preimage_of_equiv {ν : measure β} (h : is_fundamental_domain G s μ) {f : β → α} (hf : quasi_measure_preserving f ν μ) {e : G → H} (he : bijective e) @@ -134,15 +126,12 @@ h.pairwise_ae_disjoint.mono $ λ g₁ g₂ H, hν H is_fundamental_domain H (f ⁻¹' s) ν := { null_measurable_set := h.null_measurable_set.preimage hf, ae_covers := (hf.ae h.ae_covers).mono $ λ x ⟨g, hg⟩, ⟨e g, by rwa [mem_preimage, hef g x]⟩, - ae_disjoint := λ g hg, + ae_disjoint := λ a b hab, begin lift e to G ≃ H using he, - have : (e.symm g⁻¹)⁻¹ ≠ (e.symm 1)⁻¹, by simp [hg], - convert (h.pairwise_ae_disjoint this).preimage hf using 1, - { simp only [← preimage_smul_inv, preimage_preimage, ← hef _ _, e.apply_symm_apply, - inv_inv] }, - { ext1 x, - simp only [mem_preimage, ← preimage_smul, ← hef _ _, e.apply_symm_apply, one_smul] } + have : (e.symm a⁻¹)⁻¹ ≠ (e.symm b⁻¹)⁻¹, by simp [hab], + convert (h.ae_disjoint this).preimage hf using 1, + simp only [←preimage_smul_inv, preimage_preimage, ←hef _ _, e.apply_symm_apply, inv_inv], end } @[to_additive] lemma image_of_equiv {ν : measure β} (h : is_fundamental_domain G s μ) @@ -156,11 +145,9 @@ begin rw [← hef _ _, f.symm_apply_apply, f.symm_apply_apply, e.apply_symm_apply] end -@[to_additive] lemma smul (h : is_fundamental_domain G s μ) (g : G) : - is_fundamental_domain G (g • s) μ := -h.image_of_equiv (mul_action.to_perm g) (measure_preserving_smul _ _).quasi_measure_preserving - ⟨λ g', g⁻¹ * g' * g, λ g', g * g' * g⁻¹, λ g', by simp [mul_assoc], λ g', by simp [mul_assoc]⟩ $ - λ g' x, by simp [smul_smul, mul_assoc] +@[to_additive] lemma pairwise_ae_disjoint_of_ac {ν} (h : is_fundamental_domain G s μ) (hν : ν ≪ μ) : + pairwise (λ g₁ g₂ : G, ae_disjoint ν (g₁ • s) (g₂ • s)) := +h.ae_disjoint.mono $ λ g₁ g₂ H, hν H @[to_additive] lemma smul_of_comm {G' : Type*} [group G'] [mul_action G' α] [measurable_space G'] [has_measurable_smul G' α] [smul_invariant_measure G' α μ] [smul_comm_class G' G α] @@ -169,11 +156,27 @@ h.image_of_equiv (mul_action.to_perm g) (measure_preserving_smul _ _).quasi_meas h.image_of_equiv (mul_action.to_perm g) (measure_preserving_smul _ _).quasi_measure_preserving (equiv.refl _) $ smul_comm g +variables [measurable_space G] [has_measurable_smul G α] [smul_invariant_measure G α μ] + +@[to_additive] lemma null_measurable_set_smul (h : is_fundamental_domain G s μ) (g : G) : + null_measurable_set (g • s) μ := +h.null_measurable_set.smul g + +@[to_additive] lemma restrict_restrict (h : is_fundamental_domain G s μ) (g : G) (t : set α) : + (μ.restrict t).restrict (g • s) = μ.restrict (g • s ∩ t) := +restrict_restrict₀ ((h.null_measurable_set_smul g).mono restrict_le_self) + +@[to_additive] lemma smul (h : is_fundamental_domain G s μ) (g : G) : + is_fundamental_domain G (g • s) μ := +h.image_of_equiv (mul_action.to_perm g) (measure_preserving_smul _ _).quasi_measure_preserving + ⟨λ g', g⁻¹ * g' * g, λ g', g * g' * g⁻¹, λ g', by simp [mul_assoc], λ g', by simp [mul_assoc]⟩ $ + λ g' x, by simp [smul_smul, mul_assoc] + variables [countable G] {ν : measure α} @[to_additive] lemma sum_restrict_of_ac (h : is_fundamental_domain G s μ) (hν : ν ≪ μ) : sum (λ g : G, ν.restrict (g • s)) = ν := -by rw [← restrict_Union_ae (h.pairwise_ae_disjoint.mono $ λ i j h, hν h) +by rw [← restrict_Union_ae (h.ae_disjoint.mono $ λ i j h, hν h) (λ g, (h.null_measurable_set_smul g).mono_ac hν), restrict_congr_set (hν h.Union_smul_ae_eq), restrict_univ] @@ -338,7 +341,7 @@ begin by simp only [hf, hs.restrict_restrict] ... = ∫ x in ⋃ g : G, g • s, f x ∂(μ.restrict t) : (integral_Union_ae (λ g, (hs.null_measurable_set_smul g).mono_ac hac) - (hs.pairwise_ae_disjoint.mono $ λ i j h, hac h) hft.integrable.integrable_on).symm + (hs.ae_disjoint.mono $ λ i j h, hac h) hft.integrable.integrable_on).symm ... = ∫ x in t, f x ∂μ : by rw [restrict_congr_set (hac hs.Union_smul_ae_eq), restrict_univ] }, { rw [integral_undef hfs, integral_undef], From 809e920edfa343283cea507aedff916ea0f1bd88 Mon Sep 17 00:00:00 2001 From: Oliver Nash Date: Fri, 16 Dec 2022 12:14:28 +0000 Subject: [PATCH 0063/1291] feat(dynamics/ergodic/ergodic): expand ergodic map API for finite measures (#17864) --- src/dynamics/ergodic/ergodic.lean | 71 ++++++++++++++++--- src/measure_theory/measure/measure_space.lean | 24 ++++++- .../measure/measure_space_def.lean | 6 +- 3 files changed, 88 insertions(+), 13 deletions(-) diff --git a/src/dynamics/ergodic/ergodic.lean b/src/dynamics/ergodic/ergodic.lean index 0adf23a110236..99eb0e4bedfa6 100644 --- a/src/dynamics/ergodic/ergodic.lean +++ b/src/dynamics/ergodic/ergodic.lean @@ -29,6 +29,7 @@ preserving condition is relaxed to quasi measure preserving. -/ open set function filter measure_theory measure_theory.measure +open_locale ennreal variables {α : Type*} {m : measurable_space α} (f : α → α) {s : set α} include m @@ -107,15 +108,6 @@ end end measure_theory.measure_preserving -namespace ergodic - -/-- An ergodic map is quasi ergodic. -/ -lemma quasi_ergodic (hf : ergodic f μ) : quasi_ergodic f μ := -{ .. hf.to_pre_ergodic, - .. hf.to_measure_preserving.quasi_measure_preserving, } - -end ergodic - namespace quasi_ergodic /-- For a quasi ergodic map, sets that are almost invariant (rather than strictly invariant) are @@ -131,3 +123,64 @@ begin end end quasi_ergodic + +namespace ergodic + +/-- An ergodic map is quasi ergodic. -/ +lemma quasi_ergodic (hf : ergodic f μ) : quasi_ergodic f μ := +{ .. hf.to_pre_ergodic, + .. hf.to_measure_preserving.quasi_measure_preserving, } + +/-- See also `ergodic.ae_empty_or_univ_of_preimage_ae_le`. -/ +lemma ae_empty_or_univ_of_preimage_ae_le' + (hf : ergodic f μ) (hs : measurable_set s) (hs' : f⁻¹' s ≤ᵐ[μ] s) (h_fin : μ s ≠ ∞) : + s =ᵐ[μ] (∅ : set α) ∨ s =ᵐ[μ] univ := +begin + refine hf.quasi_ergodic.ae_empty_or_univ' hs _, + refine ae_eq_of_ae_subset_of_measure_ge hs' (hf.measure_preimage hs).symm.le _ h_fin, + exact measurable_set_preimage hf.measurable hs, +end + +/-- See also `ergodic.ae_empty_or_univ_of_ae_le_preimage`. -/ +lemma ae_empty_or_univ_of_ae_le_preimage' + (hf : ergodic f μ) (hs : measurable_set s) (hs' : s ≤ᵐ[μ] f⁻¹' s) (h_fin : μ s ≠ ∞) : + s =ᵐ[μ] (∅ : set α) ∨ s =ᵐ[μ] univ := +begin + replace h_fin : μ (f⁻¹' s) ≠ ∞, { rwa hf.measure_preimage hs, }, + refine hf.quasi_ergodic.ae_empty_or_univ' hs _, + exact (ae_eq_of_ae_subset_of_measure_ge hs' (hf.measure_preimage hs).le hs h_fin).symm, +end + +/-- See also `ergodic.ae_empty_or_univ_of_image_ae_le`. -/ +lemma ae_empty_or_univ_of_image_ae_le' + (hf : ergodic f μ) (hs : measurable_set s) (hs' : f '' s ≤ᵐ[μ] s) (h_fin : μ s ≠ ∞) : + s =ᵐ[μ] (∅ : set α) ∨ s =ᵐ[μ] univ := +begin + replace hs' : s ≤ᵐ[μ] f ⁻¹' s := + (has_subset.subset.eventually_le (subset_preimage_image f s)).trans + (hf.quasi_measure_preserving.preimage_mono_ae hs'), + exact ae_empty_or_univ_of_ae_le_preimage' hf hs hs' h_fin, +end + +section is_finite_measure + +variables [is_finite_measure μ] + +lemma ae_empty_or_univ_of_preimage_ae_le + (hf : ergodic f μ) (hs : measurable_set s) (hs' : f⁻¹' s ≤ᵐ[μ] s) : + s =ᵐ[μ] (∅ : set α) ∨ s =ᵐ[μ] univ := +ae_empty_or_univ_of_preimage_ae_le' hf hs hs' $ measure_ne_top μ s + +lemma ae_empty_or_univ_of_ae_le_preimage + (hf : ergodic f μ) (hs : measurable_set s) (hs' : s ≤ᵐ[μ] f⁻¹' s) : + s =ᵐ[μ] (∅ : set α) ∨ s =ᵐ[μ] univ := +ae_empty_or_univ_of_ae_le_preimage' hf hs hs' $ measure_ne_top μ s + +lemma ae_empty_or_univ_of_image_ae_le + (hf : ergodic f μ) (hs : measurable_set s) (hs' : f '' s ≤ᵐ[μ] s) : + s =ᵐ[μ] (∅ : set α) ∨ s =ᵐ[μ] univ := +ae_empty_or_univ_of_image_ae_le' hf hs hs' $ measure_ne_top μ s + +end is_finite_measure + +end ergodic diff --git a/src/measure_theory/measure/measure_space.lean b/src/measure_theory/measure/measure_space.lean index 8c561cdb00abd..bb0575a956c8c 100644 --- a/src/measure_theory/measure/measure_space.lean +++ b/src/measure_theory/measure/measure_space.lean @@ -267,12 +267,30 @@ lemma measure_eq_measure_larger_of_between_null_diff {s₁ s₂ s₃ : set α} lemma measure_compl (h₁ : measurable_set s) (h_fin : μ s ≠ ∞) : μ (sᶜ) = μ univ - μ s := by { rw compl_eq_univ_diff, exact measure_diff (subset_univ s) h₁ h_fin } +@[simp] lemma union_ae_eq_left_iff_ae_subset : (s ∪ t : set α) =ᵐ[μ] s ↔ t ≤ᵐ[μ] s := +begin + rw ae_le_set, + refine ⟨λ h, by simpa only [union_diff_left] using (ae_eq_set.mp h).1, + λ h, eventually_le_antisymm_iff.mpr + ⟨by rwa [ae_le_set, union_diff_left], has_subset.subset.eventually_le $ subset_union_left s t⟩⟩, +end + +@[simp] lemma union_ae_eq_right_iff_ae_subset : (s ∪ t : set α) =ᵐ[μ] t ↔ s ≤ᵐ[μ] t := +by rw [union_comm, union_ae_eq_left_iff_ae_subset] + +lemma ae_eq_of_ae_subset_of_measure_ge (h₁ : s ≤ᵐ[μ] t) (h₂ : μ t ≤ μ s) (hsm : measurable_set s) + (ht : μ t ≠ ∞) : s =ᵐ[μ] t := +begin + refine eventually_le_antisymm_iff.mpr ⟨h₁, ae_le_set.mpr _⟩, + replace h₂ : μ t = μ s, from h₂.antisymm (measure_mono_ae h₁), + replace ht : μ s ≠ ∞, from h₂ ▸ ht, + rw [measure_diff' t hsm ht, measure_congr (union_ae_eq_left_iff_ae_subset.mpr h₁), h₂, tsub_self], +end + /-- If `s ⊆ t`, `μ t ≤ μ s`, `μ t ≠ ∞`, and `s` is measurable, then `s =ᵐ[μ] t`. -/ lemma ae_eq_of_subset_of_measure_ge (h₁ : s ⊆ t) (h₂ : μ t ≤ μ s) (hsm : measurable_set s) (ht : μ t ≠ ∞) : s =ᵐ[μ] t := -have A : μ t = μ s, from h₂.antisymm (measure_mono h₁), -have B : μ s ≠ ∞, from A ▸ ht, -h₁.eventually_le.antisymm $ ae_le_set.2 $ by rw [measure_diff h₁ hsm B, A, tsub_self] +ae_eq_of_ae_subset_of_measure_ge (has_subset.subset.eventually_le h₁) h₂ hsm ht lemma measure_Union_congr_of_subset [countable β] {s : β → set α} {t : β → set α} (hsub : ∀ b, s b ⊆ t b) (h_le : ∀ b, μ (t b) ≤ μ (s b)) : diff --git a/src/measure_theory/measure/measure_space_def.lean b/src/measure_theory/measure/measure_space_def.lean index 1da5a90cb851b..06ca7e30abe12 100644 --- a/src/measure_theory/measure/measure_space_def.lean +++ b/src/measure_theory/measure/measure_space_def.lean @@ -367,7 +367,11 @@ lemma ae_le_set_inter {s' t' : set α} (h : s ≤ᵐ[μ] t) (h' : s' ≤ᵐ[μ] (s ∩ s' : set α) ≤ᵐ[μ] (t ∩ t' : set α) := h.inter h' -@[simp] lemma union_ae_eq_right : (s ∪ t : set α) =ᵐ[μ] t ↔ μ (s \ t) = 0 := +lemma ae_le_set_union {s' t' : set α} (h : s ≤ᵐ[μ] t) (h' : s' ≤ᵐ[μ] t') : + (s ∪ s' : set α) ≤ᵐ[μ] (t ∪ t' : set α) := +h.union h' + +lemma union_ae_eq_right : (s ∪ t : set α) =ᵐ[μ] t ↔ μ (s \ t) = 0 := by simp [eventually_le_antisymm_iff, ae_le_set, union_diff_right, diff_eq_empty.2 (set.subset_union_right _ _)] From 2c5df64ab0ef55369dc72c7c589d3af2a5bf7966 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 16 Dec 2022 14:42:59 +0000 Subject: [PATCH 0064/1291] feat(measure_theory/measure/measure_space): A `measurable_equiv` is `quasi_measure_preserving` (#17774) --- src/dynamics/ergodic/measure_preserving.lean | 7 +++++++ src/measure_theory/measure/measure_space.lean | 4 ++++ 2 files changed, 11 insertions(+) diff --git a/src/dynamics/ergodic/measure_preserving.lean b/src/dynamics/ergodic/measure_preserving.lean index a3d3eb0ca5817..751332d237ae3 100644 --- a/src/dynamics/ergodic/measure_preserving.lean +++ b/src/dynamics/ergodic/measure_preserving.lean @@ -160,4 +160,11 @@ end end measure_preserving +namespace measurable_equiv + +lemma measure_preserving_symm (μ : measure α) (e : α ≃ᵐ β) : + measure_preserving e.symm (map e μ) μ := +(e.measurable.measure_preserving μ).symm _ + +end measurable_equiv end measure_theory diff --git a/src/measure_theory/measure/measure_space.lean b/src/measure_theory/measure/measure_space.lean index bb0575a956c8c..186bcc97cedbc 100644 --- a/src/measure_theory/measure/measure_space.lean +++ b/src/measure_theory/measure/measure_space.lean @@ -3754,6 +3754,10 @@ e.measurable_embedding.restrict_map _ _ lemma map_ae (f : α ≃ᵐ β) (μ : measure α) : filter.map f μ.ae = (map f μ).ae := by { ext s, simp_rw [mem_map, mem_ae_iff, ← preimage_compl, f.map_apply] } +lemma quasi_measure_preserving_symm (μ : measure α) (e : α ≃ᵐ β) : + quasi_measure_preserving e.symm (map e μ) μ := +⟨e.symm.measurable, by rw [measure.map_map, e.symm_comp_self, measure.map_id]; measurability⟩ + end measurable_equiv From 11bb0c9152e5d14278fb0ac5e0be6d50e2c8fa05 Mon Sep 17 00:00:00 2001 From: kkytola <“kalle.kytola@aalto.fi”> Date: Fri, 16 Dec 2022 20:19:05 +0000 Subject: [PATCH 0065/1291] feat(measure_theory/integral/layercake): Add versions of layercake formula with strict inequality. (#17920) This PR adds versions of layercake formula with measures of sets strictly above a given level, i.e., of the form `{x | f x > t}`. This is useful particularly when `f` is continuous and one wants the set in question to be open (rather than closed). Co-authored-by: Kalle-swift Co-authored-by: kkytola <39528102+kkytola@users.noreply.github.com> --- src/measure_theory/integral/layercake.lean | 154 ++++++++++++++++-- src/measure_theory/measure/measure_space.lean | 11 ++ 2 files changed, 154 insertions(+), 11 deletions(-) diff --git a/src/measure_theory/integral/layercake.lean b/src/measure_theory/integral/layercake.lean index beeef7164966a..d4975c79d6bd8 100644 --- a/src/measure_theory/integral/layercake.lean +++ b/src/measure_theory/integral/layercake.lean @@ -30,12 +30,22 @@ We also give the two most common applications of the layer cake formula * a representation of the integral of the p:th power of a nonnegative function f: ∫ f(ω)^p ∂μ(ω) = p * ∫ t^(p-1) * μ {ω | f(ω) ≥ t} dt . +Variants of the formulas with measures of sets of the form {ω | f(ω) > t} instead of {ω | f(ω) ≥ t} +are also included. + ## Main results - * `lintegral_comp_eq_lintegral_meas_le_mul`: The general layer cake formula with Lebesgue - integrals. - * `lintegral_eq_lintegral_meas_le`: The most common special case of the layer cake formula, which - states that for a nonnegative function f we have ∫ f(ω) ∂μ(ω) = ∫ μ {ω | f(ω) ≥ t} dt . + * `lintegral_comp_eq_lintegral_meas_le_mul` and `lintegral_comp_eq_lintegral_meas_lt_mul`: + The general layer cake formulas with Lebesgue integrals, written in terms of measures of + sets of the forms {ω | t ≤ f(ω)} and {ω | t < f(ω)}, respectively. + * `lintegral_eq_lintegral_meas_le` and `lintegral_eq_lintegral_meas_lt`: + The most common special cases of the layer cake formulas, stating that for a nonnegative + function f we have ∫ f(ω) ∂μ(ω) = ∫ μ {ω | f(ω) ≥ t} dt and + ∫ f(ω) ∂μ(ω) = ∫ μ {ω | f(ω) > t} dt, respectively. + * `lintegral_rpow_eq_lintegral_meas_le_mul` and `lintegral_rpow_eq_lintegral_meas_lt_mul`: + Other common special cases of the layer cake formulas, stating that for a nonnegative function f + and p > 0, we have ∫ f(ω)^p ∂μ(ω) = p * ∫ μ {ω | f(ω) ≥ t} * t^(p-1) dt and + ∫ f(ω)^p ∂μ(ω) = p * ∫ μ {ω | f(ω) > t} * t^(p-1) dt, respectively. ## Tags @@ -46,18 +56,20 @@ noncomputable theory open_locale ennreal measure_theory open set measure_theory filter -/-! ### Layercake theorem -/ +/-! ### Layercake formula -/ section layercake namespace measure_theory variables {α : Type*} [measurable_space α] {f : α → ℝ} {g : ℝ → ℝ} {s : set α} -/-- An auxiliary version of the layer cake theorem (Cavalieri's principle, tail probability +/-- An auxiliary version of the layer cake formula (Cavalieri's principle, tail probability formula), with a measurability assumption that would also essentially follow from the integrability assumptions. -See `measure_theory.layercake` for the main formulation of the layer cake theorem. -/ +See `measure_theory.lintegral_comp_eq_lintegral_meas_le_mul` and +`measure_theory.lintegral_comp_eq_lintegral_meas_lt_mul` for the main formulations of the layer +cake formula. -/ lemma lintegral_comp_eq_lintegral_meas_le_mul_of_measurable (μ : measure α) [sigma_finite μ] (f_nn : 0 ≤ f) (f_mble : measurable f) (g_intble : ∀ t > 0, interval_integrable g volume 0 t) @@ -132,7 +144,7 @@ begin exact (ennreal.measurable_of_real.comp (g_mble.comp measurable_snd)).ae_measurable.indicator mble, end -/-- The layer cake theorem / Cavalieri's principle / tail probability formula: +/-- The layer cake formula / Cavalieri's principle / tail probability formula: Let `f` be a non-negative measurable function on a sigma-finite measure space. Let `G` be an increasing absolutely continuous function on the positive real line, vanishing at the origin, @@ -140,7 +152,10 @@ with derivative `G' = g`. Then the integral of the composition `G ∘ f` can be the integral over the positive real line of the "tail measures" `μ {ω | f(ω) ≥ t}` of `f` weighted by `g`. -Roughly speaking, the statement is: `∫⁻ (G ∘ f) ∂μ = ∫⁻ t in 0 .. ∞, g(t) * μ {ω | f(ω) ≥ t}`. -/ +Roughly speaking, the statement is: `∫⁻ (G ∘ f) ∂μ = ∫⁻ t in 0 .. ∞, g(t) * μ {ω | f(ω) ≥ t}`. + +See `lintegral_comp_eq_lintegral_meas_lt_mul` for a version with sets of the form `{ω | f(ω) > t}` +instead. -/ theorem lintegral_comp_eq_lintegral_meas_le_mul (μ : measure α) [sigma_finite μ] (f_nn : 0 ≤ f) (f_mble : measurable f) (g_intble : ∀ t > 0, interval_integrable g volume 0 t) @@ -177,7 +192,10 @@ end /-- The standard case of the layer cake formula / Cavalieri's principle / tail probability formula: For a nonnegative function `f` on a sigma-finite measure space, the Lebesgue integral of `f` can -be written (roughly speaking) as: `∫⁻ f ∂μ = ∫⁻ t in 0 .. ∞, μ {ω | f(ω) ≥ t}`. -/ +be written (roughly speaking) as: `∫⁻ f ∂μ = ∫⁻ t in 0 .. ∞, μ {ω | f(ω) ≥ t}`. + +See `lintegral_eq_lintegral_meas_lt` for a version with sets of the form `{ω | f(ω) > t}` +instead. -/ theorem lintegral_eq_lintegral_meas_le (μ : measure α) [sigma_finite μ] (f_nn : 0 ≤ f) (f_mble : measurable f) : ∫⁻ ω, ennreal.of_real (f ω) ∂μ = ∫⁻ t in Ioi 0, (μ {a : α | t ≤ f a}) := @@ -196,7 +214,10 @@ end /-- An application of the layer cake formula / Cavalieri's principle / tail probability formula: For a nonnegative function `f` on a sigma-finite measure space, the Lebesgue integral of `f` can -be written (roughly speaking) as: `∫⁻ f^p ∂μ = p * ∫⁻ t in 0 .. ∞, t^(p-1) * μ {ω | f(ω) ≥ t}`. -/ +be written (roughly speaking) as: `∫⁻ f^p ∂μ = p * ∫⁻ t in 0 .. ∞, t^(p-1) * μ {ω | f(ω) ≥ t}`. + +See `lintegral_rpow_eq_lintegral_meas_lt_mul` for a version with sets of the form `{ω | f(ω) > t}` +instead. -/ theorem lintegral_rpow_eq_lintegral_meas_le_mul (μ : measure α) [sigma_finite μ] (f_nn : 0 ≤ f) (f_mble : measurable f) {p : ℝ} (p_pos: 0 < p) : ∫⁻ ω, ennreal.of_real ((f ω)^p) ∂μ @@ -226,3 +247,114 @@ end end measure_theory end layercake + +section layercake_lt + +open measure_theory + +variables {α : Type*} [measurable_space α] (μ : measure α) +variables {β : Type*} [measurable_space β] [measurable_singleton_class β] + +namespace measure + +lemma meas_le_ne_meas_lt_subset_meas_pos {R : Type*} [linear_order R] + [measurable_space R] [measurable_singleton_class R] {g : α → R} (g_mble : measurable g) {t : R} + (ht : μ {a : α | t ≤ g a} ≠ μ {a : α | t < g a}) : + 0 < μ {a : α | g a = t} := +begin + have uni : {a : α | t ≤ g a } = {a : α | t < g a} ∪ {a : α | t = g a}, + { ext a, + simp only [mem_set_of_eq, mem_union], + apply le_iff_lt_or_eq, }, + rw (show {a : α | t = g a} = {a : α | g a = t}, by simp_rw [eq_comm]) at uni, + have disj : {a : α | t < g a} ∩ {a : α | g a = t} = ∅, + { ext a, + simp only [mem_inter_iff, mem_set_of_eq, mem_empty_iff_false, iff_false, not_and], + exact ne_of_gt, }, + have μ_add : μ {a : α | t ≤ g a} = μ {a : α | t < g a} + μ {a : α | g a = t}, + by rw [uni, measure_union (disjoint_iff_inter_eq_empty.mpr disj) + (g_mble (finite.measurable_set (finite_singleton t)))], + by_contra con, + rw [not_lt, nonpos_iff_eq_zero] at con, + rw [con, add_zero] at μ_add, + exact ht μ_add, +end + +lemma countable_meas_le_ne_meas_lt [sigma_finite μ] {R : Type*} [linear_order R] + [measurable_space R] [measurable_singleton_class R] {g : α → R} (g_mble : measurable g) : + {t : R | μ {a : α | t ≤ g a } ≠ μ {a : α | t < g a}}.countable := +countable.mono (show _, from λ t ht, meas_le_ne_meas_lt_subset_meas_pos μ g_mble ht) + (measure.countable_meas_level_set_pos g_mble) + +lemma meas_le_ae_eq_meas_lt [sigma_finite μ] {R : Type*} [linear_order R] [measurable_space R] + [measurable_singleton_class R] (ν : measure R) [has_no_atoms ν] + {g : α → R} (g_mble : measurable g) : + (λ t, μ {a : α | t ≤ g a}) =ᵐ[ν] (λ t, μ {a : α | t < g a}) := +set.countable.measure_zero (measure.countable_meas_le_ne_meas_lt μ g_mble) _ + +end measure + +variables {f : α → ℝ} {g : ℝ → ℝ} {s : set α} + +/-- The layer cake formula / Cavalieri's principle / tail probability formula: + +Let `f` be a non-negative measurable function on a sigma-finite measure space. Let `G` be an +increasing absolutely continuous function on the positive real line, vanishing at the origin, +with derivative `G' = g`. Then the integral of the composition `G ∘ f` can be written as +the integral over the positive real line of the "tail measures" `μ {ω | f(ω) > t}` of `f` +weighted by `g`. + +Roughly speaking, the statement is: `∫⁻ (G ∘ f) ∂μ = ∫⁻ t in 0 .. ∞, g(t) * μ {ω | f(ω) > t}`. + +See `lintegral_comp_eq_lintegral_meas_le_mul` for a version with sets of the form `{ω | f(ω) ≥ t}` +instead. -/ +theorem lintegral_comp_eq_lintegral_meas_lt_mul (μ : measure α) [sigma_finite μ] + (f_nn : 0 ≤ f) (f_mble : measurable f) + (g_intble : ∀ t > 0, interval_integrable g volume 0 t) + (g_nn : ∀ᵐ t ∂(volume.restrict (Ioi 0)), 0 ≤ g t) : + ∫⁻ ω, ennreal.of_real (∫ t in 0 .. f ω, g t) ∂μ + = ∫⁻ t in Ioi 0, μ {a : α | t < f a} * ennreal.of_real (g t) := +begin + rw lintegral_comp_eq_lintegral_meas_le_mul μ f_nn f_mble g_intble g_nn, + apply lintegral_congr_ae, + filter_upwards [measure.meas_le_ae_eq_meas_lt μ (volume.restrict (Ioi 0)) f_mble] with t ht, + rw ht, +end + +/-- The standard case of the layer cake formula / Cavalieri's principle / tail probability formula: + +For a nonnegative function `f` on a sigma-finite measure space, the Lebesgue integral of `f` can +be written (roughly speaking) as: `∫⁻ f ∂μ = ∫⁻ t in 0 .. ∞, μ {ω | f(ω) > t}`. + +See `lintegral_eq_lintegral_meas_le` for a version with sets of the form `{ω | f(ω) ≥ t}` +instead. -/ +theorem lintegral_eq_lintegral_meas_lt (μ : measure α) [sigma_finite μ] + (f_nn : 0 ≤ f) (f_mble : measurable f) : + ∫⁻ ω, ennreal.of_real (f ω) ∂μ = ∫⁻ t in Ioi 0, (μ {a : α | t < f a}) := +begin + rw lintegral_eq_lintegral_meas_le μ f_nn f_mble, + apply lintegral_congr_ae, + filter_upwards [measure.meas_le_ae_eq_meas_lt μ (volume.restrict (Ioi 0)) f_mble] with t ht, + rw ht, +end + +/-- An application of the layer cake formula / Cavalieri's principle / tail probability formula: + +For a nonnegative function `f` on a sigma-finite measure space, the Lebesgue integral of `f` can +be written (roughly speaking) as: `∫⁻ f^p ∂μ = p * ∫⁻ t in 0 .. ∞, t^(p-1) * μ {ω | f(ω) > t}`. + +See `lintegral_rpow_eq_lintegral_meas_le_mul` for a version with sets of the form `{ω | f(ω) ≥ t}` +instead. -/ +theorem lintegral_rpow_eq_lintegral_meas_lt_mul (μ : measure α) [sigma_finite μ] + (f_nn : 0 ≤ f) (f_mble : measurable f) {p : ℝ} (p_pos: 0 < p) : + ∫⁻ ω, ennreal.of_real ((f ω)^p) ∂μ + = (ennreal.of_real p) * ∫⁻ t in Ioi 0, (μ {a : α | t < f a}) * ennreal.of_real (t^(p-1)) := +begin + rw lintegral_rpow_eq_lintegral_meas_le_mul μ f_nn f_mble p_pos, + apply congr_arg (λ z, (ennreal.of_real p * z)), + apply lintegral_congr_ae, + filter_upwards [measure.meas_le_ae_eq_meas_lt μ (volume.restrict (Ioi 0)) f_mble] with t ht, + rw ht, +end + +end layercake_lt diff --git a/src/measure_theory/measure/measure_space.lean b/src/measure_theory/measure/measure_space.lean index 186bcc97cedbc..99c74342f3dc9 100644 --- a/src/measure_theory/measure/measure_space.lean +++ b/src/measure_theory/measure/measure_space.lean @@ -3134,6 +3134,17 @@ begin exact Union_subset (λ i, inter_subset_right _ _), }, end +lemma countable_meas_level_set_pos {α β : Type*} + [measurable_space α] {μ : measure α} [sigma_finite μ] + [measurable_space β] [measurable_singleton_class β] {g : α → β} (g_mble : measurable g) : + set.countable {t : β | 0 < μ {a : α | g a = t}} := +begin + have level_sets_disjoint : pairwise (disjoint on (λ (t : β), {a : α | g a = t})), + from λ s t hst, disjoint.preimage g (disjoint_singleton.mpr hst), + exact measure.countable_meas_pos_of_disjoint_Union + (λ b, g_mble (‹measurable_singleton_class β›.measurable_set_singleton b)) level_sets_disjoint, +end + /-- The measurable superset `to_measurable μ t` of `t` (which has the same measure as `t`) satisfies, for any measurable set `s`, the equality `μ (to_measurable μ t ∩ s) = μ (t ∩ s)`. This only holds when `μ` is σ-finite. For a version without this assumption (but requiring From 706d88f2b8fdfeb0b22796433d7a6c1a010af9f2 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9mi=20Bottinelli?= Date: Sat, 17 Dec 2022 06:45:00 +0000 Subject: [PATCH 0066/1291] feat(combinatorics/quiver/*): More edge reversal constructions (#17665) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Move * `quiver.symmetrify`, * `quiver.has_reverse`, * `quiver.has_involutive_reverse`, * `quiver.reverse`, * `quiver.path.reverse`, * `quiver.symmetrify.of`, * `quiver.lift`, * associated lemmas, from `combinatorics/quiver/connected_component.lean` to `combinatorics/quiver/symmetrify.lean`. Add * the class `prefunctor.map_reverse` witnessing that a prefunctor maps reverses to reverses, and change the lemmas taking this property as an explicit argument. * `prefunctor.symmetrify` mapping a prefunctor to the prefunctor between symmetrifications, * associated lemmas, to `combinatorics/quiver/symmetrify.lean`. Move `quiver.hom.to_pos` and `quiver.hom.to_neg` from `category_theory/groupoid/free_groupoid.lean` to `combinatorics/quiver/symmetrify.lean`. Add `map_reverse` instance for functors between groupoids. Co-authored-by: Rémi Bottinelli --- src/category_theory/groupoid.lean | 6 + .../groupoid/free_groupoid.lean | 13 +- .../quiver/connected_component.lean | 117 +--------- src/combinatorics/quiver/push.lean | 4 + src/combinatorics/quiver/symmetric.lean | 215 ++++++++++++++++++ 5 files changed, 235 insertions(+), 120 deletions(-) create mode 100644 src/combinatorics/quiver/symmetric.lean diff --git a/src/category_theory/groupoid.lean b/src/category_theory/groupoid.lean index 34075a9674b96..c3b0ca0a31ab6 100644 --- a/src/category_theory/groupoid.lean +++ b/src/category_theory/groupoid.lean @@ -75,6 +75,12 @@ instance groupoid_has_involutive_reverse : quiver.has_involutive_reverse C := @[simp] lemma groupoid.reverse_eq_inv (f : X ⟶ Y) : quiver.reverse f = groupoid.inv f := rfl +instance functor_map_reverse {D : Type*} [groupoid D] (F : C ⥤ D) : + F.to_prefunctor.map_reverse := +{ map_reverse' := λ X Y f, by + simp only [quiver.reverse, quiver.has_reverse.reverse', groupoid.inv_eq_inv, + functor.to_prefunctor_map, functor.map_inv], } + variables (X Y) /-- In a groupoid, isomorphisms are equivalent to morphisms. -/ diff --git a/src/category_theory/groupoid/free_groupoid.lean b/src/category_theory/groupoid/free_groupoid.lean index b3b2c8ff30833..4a683ef2d8107 100644 --- a/src/category_theory/groupoid/free_groupoid.lean +++ b/src/category_theory/groupoid/free_groupoid.lean @@ -9,6 +9,7 @@ import category_theory.groupoid import tactic.nth_rewrite import category_theory.path_category import category_theory.quotient +import combinatorics.quiver.symmetric /-! # Free groupoid on a quiver @@ -45,14 +46,6 @@ universes u v u' v' u'' v'' variables {V : Type u} [quiver.{v+1} V] -/-- Shorthand for the "forward" arrow corresponding to `f` in `symmetrify V` -/ -abbreviation quiver.hom.to_pos {X Y : V} (f : X ⟶ Y) : - (quiver.symmetrify_quiver V).hom X Y := sum.inl f - -/-- Shorthand for the "backward" arrow corresponding to `f` in `symmetrify V` -/ -abbreviation quiver.hom.to_neg {X Y : V} (f : X ⟶ Y) : - (quiver.symmetrify_quiver V).hom Y X := sum.inr f - /-- Shorthand for the "forward" arrow corresponding to `f` in `paths $ symmetrify V` -/ abbreviation quiver.hom.to_pos_path {X Y : V} (f : X ⟶ Y) : ((category_theory.paths.category_paths $ quiver.symmetrify V).hom X Y) := f.to_pos.to_path @@ -168,9 +161,9 @@ lemma lift_unique (φ : V ⥤q V') (Φ : free_groupoid V ⥤ V') begin apply quotient.lift_unique, apply paths.lift_unique, - apply quiver.symmetrify.lift_unique, + fapply @quiver.symmetrify.lift_unique _ _ _ _ _ _ _ _ _, { rw ←functor.to_prefunctor_comp, exact hΦ, }, - { rintros X Y f, + { constructor, rintros X Y f, simp only [←functor.to_prefunctor_comp,prefunctor.comp_map, paths.of_map, inv_eq_inv], change Φ.map (inv ((quotient.functor red_step).to_prefunctor.map f.to_path)) = inv (Φ.map ((quotient.functor red_step).to_prefunctor.map f.to_path)), diff --git a/src/combinatorics/quiver/connected_component.lean b/src/combinatorics/quiver/connected_component.lean index 303b8dcee109b..1a27252fabbaa 100644 --- a/src/combinatorics/quiver/connected_component.lean +++ b/src/combinatorics/quiver/connected_component.lean @@ -5,8 +5,7 @@ Authors: David Wärn -/ import combinatorics.quiver.subquiver import combinatorics.quiver.path -import data.sum.basic - +import combinatorics.quiver.symmetric /-! ## Weakly connected components @@ -14,120 +13,18 @@ import data.sum.basic > https://github.com/leanprover-community/mathlib4/pull/836 > Any changes to this file require a corresponding PR to mathlib4. -For a quiver `V`, we build a quiver `symmetrify V` by adding a reversal of every edge. -Informally, a path in `symmetrify V` corresponds to a 'zigzag' in `V`. This lets us -define the type `weakly_connected_component V` as the quotient of `V` by the relation which -identifies `a` with `b` if there is a path from `a` to `b` in `symmetrify V`. (These -zigzags can be seen as a proof-relevant analogue of `eqv_gen`.) + +For a quiver `V`, we define the type `weakly_connected_component V` as the quotient of `V` +by the relation which identifies `a` with `b` if there is a path from `a` to `b` in `symmetrify V`. +(These zigzags can be seen as a proof-relevant analogue of `eqv_gen`.) Strongly connected components have not yet been defined. -/ -universes v u +universes u namespace quiver -/-- A type synonym for the symmetrized quiver (with an arrow both ways for each original arrow). - NB: this does not work for `Prop`-valued quivers. It requires `[quiver.{v+1} V]`. -/ -@[nolint has_nonempty_instance] -def symmetrify (V) : Type u := V - -instance symmetrify_quiver (V : Type u) [quiver V] : quiver (symmetrify V) := -⟨λ a b : V, (a ⟶ b) ⊕ (b ⟶ a)⟩ - -variables (V : Type u) [quiver.{v+1} V] - -/-- A quiver `has_reverse` if we can reverse an arrow `p` from `a` to `b` to get an arrow - `p.reverse` from `b` to `a`.-/ -class has_reverse := -(reverse' : Π {a b : V}, (a ⟶ b) → (b ⟶ a)) - -/-- Reverse the direction of an arrow. -/ -def reverse {V} [quiver.{v+1} V] [has_reverse V] {a b : V} : (a ⟶ b) → (b ⟶ a) := - has_reverse.reverse' - -/-- A quiver `has_involutive_reverse` if reversing twice is the identity.`-/ -class has_involutive_reverse extends has_reverse V := -(inv' : Π {a b : V} (f : a ⟶ b), reverse (reverse f) = f) - -@[simp] lemma reverse_reverse {V} [quiver.{v+1} V] [h : has_involutive_reverse V] - {a b : V} (f : a ⟶ b) : reverse (reverse f) = f := by apply h.inv' - -variables {V} - -instance : has_reverse (symmetrify V) := ⟨λ a b e, e.swap⟩ -instance : has_involutive_reverse (symmetrify V) := -{ to_has_reverse := ⟨λ a b e, e.swap⟩, - inv' := λ a b e, congr_fun sum.swap_swap_eq e } - - -/-- Reverse the direction of a path. -/ -@[simp] def path.reverse [has_reverse V] {a : V} : Π {b}, path a b → path b a -| a path.nil := path.nil -| b (path.cons p e) := (reverse e).to_path.comp p.reverse - -@[simp] lemma path.reverse_to_path [has_reverse V] {a b : V} (f : a ⟶ b) : - f.to_path.reverse = (reverse f).to_path := rfl - -@[simp] lemma path.reverse_comp [has_reverse V] {a b c : V} (p : path a b) (q : path b c) : - (p.comp q).reverse = q.reverse.comp p.reverse := by -{ induction q, { simp, }, { simp [q_ih], }, } - -@[simp] lemma path.reverse_reverse [h : has_involutive_reverse V] {a b : V} (p : path a b) : - p.reverse.reverse = p := by -{ induction p, - { simp, }, - { simp only [path.reverse, path.reverse_comp, path.reverse_to_path, reverse_reverse, p_ih], - refl, }, } - -/-- The inclusion of a quiver in its symmetrification -/ -def symmetrify.of : prefunctor V (symmetrify V) := -{ obj := id, - map := λ X Y f, sum.inl f } - -/-- Given a quiver `V'` with reversible arrows, a prefunctor to `V'` can be lifted to one from - `symmetrify V` to `V'` -/ -def symmetrify.lift {V' : Type*} [quiver V'] [has_reverse V'] (φ : prefunctor V V') : - prefunctor (symmetrify V) V' := -{ obj := φ.obj, - map := λ X Y f, sum.rec (λ fwd, φ.map fwd) (λ bwd, reverse (φ.map bwd)) f } - -lemma symmetrify.lift_spec (V' : Type*) [quiver V'] [has_reverse V'] (φ : prefunctor V V') : - symmetrify.of.comp (symmetrify.lift φ) = φ := -begin - fapply prefunctor.ext, - { rintro X, refl, }, - { rintros X Y f, refl, }, -end - -lemma symmetrify.lift_reverse (V' : Type*) [quiver V'] [h : has_involutive_reverse V'] - (φ : prefunctor V V') - {X Y : symmetrify V} (f : X ⟶ Y) : - (symmetrify.lift φ).map (quiver.reverse f) = quiver.reverse ((symmetrify.lift φ).map f) := -begin - dsimp [symmetrify.lift], cases f, - { simp only, refl, }, - { simp only, rw h.inv', refl, } -end - -/-- `lift φ` is the only prefunctor extending `φ` and preserving reverses. -/ -lemma symmetrify.lift_unique (V' : Type*) [quiver V'] [has_reverse V'] - (φ : prefunctor V V') - (Φ : prefunctor (symmetrify V) V') - (hΦ : symmetrify.of.comp Φ = φ) - (hΦinv : ∀ {X Y : V} (f : X ⟶ Y), Φ.map (reverse f) = reverse (Φ.map f)) : - Φ = symmetrify.lift φ := -begin - subst_vars, - fapply prefunctor.ext, - { rintro X, refl, }, - { rintros X Y f, - cases f, - { refl, }, - { dsimp [symmetrify.lift,symmetrify.of], - convert hΦinv (sum.inl f), }, }, -end - -variables (V) +variables (V : Type*) [quiver.{u+1} V] /-- Two vertices are related in the zigzag setoid if there is a zigzag of arrows from one to the other. -/ diff --git a/src/combinatorics/quiver/push.lean b/src/combinatorics/quiver/push.lean index 0a373bd18b602..6fc5366b12342 100644 --- a/src/combinatorics/quiver/push.lean +++ b/src/combinatorics/quiver/push.lean @@ -21,6 +21,8 @@ universes v v₁ v₂ u u₁ u₂ variables {V : Type*} [quiver V] {W : Type*} (σ : V → W) +namespace quiver + /-- The `quiver` instance obtained by pushing arrows of `V` along the map `σ : V → W` -/ @[nolint unused_arguments] def push (σ : V → W) := W @@ -80,3 +82,5 @@ begin end end push + +end quiver diff --git a/src/combinatorics/quiver/symmetric.lean b/src/combinatorics/quiver/symmetric.lean new file mode 100644 index 0000000000000..156867f50782e --- /dev/null +++ b/src/combinatorics/quiver/symmetric.lean @@ -0,0 +1,215 @@ +/- +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 +-/ +import combinatorics.quiver.basic +import combinatorics.quiver.path +import combinatorics.quiver.push +import data.sum.basic +/-! +## Symmetric quivers and arrow reversal + +This file contains constructions related to symmetric quivers: + +* `symmetrify V` adds formal inverses to each arrow of `V`. +* `has_reverse` is the class of quivers where each arrow has an assigned formal inverse. +* `has_involutive_reverse` extends `has_reverse` by requiring that the reverse of the reverse + is equal to the original arrow. +* `prefunctor.preserve_reverse` is the class of prefunctors mapping reverses to reverses. +* `symmetrify.of`, `symmetrify.lift`, and the associated lemmas witness the universal property + of `symmetrify`. +-/ + +universes v u w v' + +namespace quiver + +/-- A type synonym for the symmetrized quiver (with an arrow both ways for each original arrow). + NB: this does not work for `Prop`-valued quivers. It requires `[quiver.{v+1} V]`. -/ +@[nolint has_nonempty_instance] +def symmetrify (V : Type*) := V + +instance symmetrify_quiver (V : Type u) [quiver V] : quiver (symmetrify V) := +⟨λ a b : V, (a ⟶ b) ⊕ (b ⟶ a)⟩ + +variables (U V W : Type*) [quiver.{u+1} U] [quiver.{v+1} V] [quiver.{w+1} W] + +/-- A quiver `has_reverse` if we can reverse an arrow `p` from `a` to `b` to get an arrow + `p.reverse` from `b` to `a`.-/ +class has_reverse := +(reverse' : Π {a b : V}, (a ⟶ b) → (b ⟶ a)) + +/-- Reverse the direction of an arrow. -/ +def reverse {V} [quiver.{v+1} V] [has_reverse V] {a b : V} : (a ⟶ b) → (b ⟶ a) := +has_reverse.reverse' + +/-- A quiver `has_involutive_reverse` if reversing twice is the identity.`-/ +class has_involutive_reverse extends has_reverse V := +(inv' : Π {a b : V} (f : a ⟶ b), reverse (reverse f) = f) + +variables {U V W} + +@[simp] lemma reverse_reverse [h : has_involutive_reverse V] + {a b : V} (f : a ⟶ b) : reverse (reverse f) = f := h.inv' f + +@[simp] lemma reverse_inj [has_involutive_reverse V] + {a b : V} (f g : a ⟶ b) : reverse f = reverse g ↔ f = g := +begin + split, + { rintro h, simpa using congr_arg quiver.reverse h, }, + { rintro h, congr, assumption, }, +end + +lemma eq_reverse_iff [has_involutive_reverse V] + {a b : V} (f : a ⟶ b) (g : b ⟶ a) : f = reverse g ↔ reverse f = g := +by rw [←reverse_inj, reverse_reverse] + +section map_reverse + +variables [has_reverse U] [has_reverse V] [has_reverse W] + +/-- A prefunctor preserving reversal of arrows -/ +class _root_.prefunctor.map_reverse (φ : U ⥤q V) := +(map_reverse' : ∀ {u v : U} (e : u ⟶ v), φ.map (reverse e) = reverse (φ.map e)) + +@[simp] lemma _root_.prefunctor.map_reverse' (φ : U ⥤q V) [φ.map_reverse] {u v : U} (e : u ⟶ v) : + φ.map (reverse e) = reverse (φ.map e) := +prefunctor.map_reverse.map_reverse' e + +instance _root_.prefunctor.map_reverse_comp (φ : U ⥤q V) (ψ : V ⥤q W) + [φ.map_reverse] [ψ.map_reverse] : (φ ⋙q ψ).map_reverse := +{ map_reverse' := λ u v e, by { simp only [prefunctor.comp_map, prefunctor.map_reverse'], } } + +instance _root_.prefunctor.map_reverse_id : (prefunctor.id U).map_reverse := +{ map_reverse' := λ u v e, rfl } + +end map_reverse + +instance : has_reverse (symmetrify V) := ⟨λ a b e, e.swap⟩ +instance : has_involutive_reverse (symmetrify V) := +{ reverse' := λ _ _ e, e.swap, + inv' := λ _ _ e, congr_fun sum.swap_swap_eq e } + +@[simp] lemma symmetrify_reverse {a b : symmetrify V} (e : a ⟶ b) : + reverse e = e.swap := rfl + +/-- Shorthand for the "forward" arrow corresponding to `f` in `symmetrify V` -/ +abbreviation hom.to_pos {X Y : V} (f : X ⟶ Y) : + (quiver.symmetrify_quiver V).hom X Y := sum.inl f + +/-- Shorthand for the "backward" arrow corresponding to `f` in `symmetrify V` -/ +abbreviation hom.to_neg {X Y : V} (f : X ⟶ Y) : + (quiver.symmetrify_quiver V).hom Y X := sum.inr f + +/-- Reverse the direction of a path. -/ +@[simp] def path.reverse [has_reverse V] {a : V} : Π {b}, path a b → path b a +| a path.nil := path.nil +| b (path.cons p e) := (reverse e).to_path.comp p.reverse + +@[simp] lemma path.reverse_to_path [has_reverse V] {a b : V} (f : a ⟶ b) : + f.to_path.reverse = (reverse f).to_path := rfl + +@[simp] lemma path.reverse_comp [has_reverse V] {a b c : V} (p : path a b) (q : path b c) : + (p.comp q).reverse = q.reverse.comp p.reverse := by +{ induction q, { simp, }, { simp [q_ih], }, } + +@[simp] lemma path.reverse_reverse [has_involutive_reverse V] {a b : V} (p : path a b) : + p.reverse.reverse = p := +begin + induction p, + { simp, }, + { simp only [path.reverse, path.reverse_comp, path.reverse_to_path, reverse_reverse, p_ih], + refl, }, +end + +namespace symmetrify + +/-- The inclusion of a quiver in its symmetrification -/ +@[simps] def of : V ⥤q symmetrify V := +{ obj := id, + map := λ X Y f, sum.inl f } + +variables {V' : Type*} [quiver.{v'+1} V'] + +/-- Given a quiver `V'` with reversible arrows, a prefunctor to `V'` can be lifted to one from + `symmetrify V` to `V'` -/ +def lift [has_reverse V'] (φ : V ⥤q V') : + (symmetrify V) ⥤q V' := +{ obj := φ.obj, + map := λ X Y f, sum.rec (λ fwd, φ.map fwd) (λ bwd, reverse (φ.map bwd)) f } + +lemma lift_spec [has_reverse V'] (φ : V ⥤q V') : + of ⋙q (lift φ) = φ := +begin + fapply prefunctor.ext, + { rintro X, refl, }, + { rintros X Y f, refl, }, +end + +lemma lift_reverse [h : has_involutive_reverse V'] + (φ : V ⥤q V') + {X Y : symmetrify V} (f : X ⟶ Y) : + (lift φ).map (quiver.reverse f) = quiver.reverse ((lift φ).map f) := +begin + dsimp [lift], cases f, + { simp only, refl, }, + { simp only [reverse_reverse], refl, } +end + +/-- `lift φ` is the only prefunctor extending `φ` and preserving reverses. -/ +lemma lift_unique [has_reverse V'] + (φ : V ⥤q V') + (Φ : (symmetrify V) ⥤q V') + (hΦ : of ⋙q Φ = φ) [hΦrev : Φ.map_reverse] : + Φ = lift φ := +begin + subst_vars, + fapply prefunctor.ext, + { rintro X, refl, }, + { rintros X Y f, + cases f, + { refl, }, + { dsimp [lift,of], + simp only [←prefunctor.map_reverse', symmetrify_reverse, sum.swap_inl], }, }, +end + +/-- A prefunctor canonically defines a prefunctor of the symmetrifications. -/ +@[simps] def _root_.prefunctor.symmetrify (φ : U ⥤q V) : (symmetrify U) ⥤q (symmetrify V) := +{ obj := φ.obj, + map := λ X Y, sum.map φ.map φ.map } + +instance _root_.prefunctor.symmetrify_map_reverse (φ : U ⥤q V) : + prefunctor.map_reverse φ.symmetrify := ⟨λ u v e, by { cases e; refl }⟩ + +end symmetrify + +namespace push + +variables {V' : Type*} (σ : V → V') + +instance [has_reverse V] : has_reverse (push σ) := +{ reverse' := λ a b F, by { cases F, constructor, apply reverse, exact F_f, } } + +instance [has_involutive_reverse V] : has_involutive_reverse (push σ) := +{ reverse' := λ a b F, by { cases F, constructor, apply reverse, exact F_f, }, + inv' := λ a b F, by { cases F, dsimp [reverse], congr, apply reverse_reverse, } } + +lemma of_reverse [h : has_involutive_reverse V] (X Y : V) (f : X ⟶ Y): + (reverse $ ((push.of σ)).map f) = ((push.of σ)).map (reverse f) := rfl + +instance of_map_reverse [h : has_involutive_reverse V] : (push.of σ).map_reverse := +⟨ by simp [of_reverse] ⟩ + +end push + +/-- +A quiver is preconnected iff there exists a path between any pair of +vertices. +Note that if `V` doesn't `has_reverse`, then the definition is stronger than +simply having a preconnected underlying `simple_graph`, since a path in one +direction doesn't induce one in the other. +-/ +def is_preconnected (V) [quiver.{u+1} V] := ∀ (X Y : V), nonempty (path X Y) + +end quiver From c2337ec8cf5e6cbd67dfe006079c9b31e224bf85 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sat, 17 Dec 2022 09:24:27 +0000 Subject: [PATCH 0067/1291] chore(number_theory/divisors): golf (#17954) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit * Upgrade `nat.swap_mem_divisors_antidiagonal` to a `simp` `iff` lemma. * Use `(equiv.prod_comm _ _).to_embedding` instead of `⟨prod.swap, _⟩`. * Golf. --- src/number_theory/divisors.lean | 45 +++++++++++---------------------- 1 file changed, 15 insertions(+), 30 deletions(-) diff --git a/src/number_theory/divisors.lean b/src/number_theory/divisors.lean index 9cbb0d1fed381..e239857c3e6ad 100644 --- a/src/number_theory/divisors.lean +++ b/src/number_theory/divisors.lean @@ -181,12 +181,9 @@ lemma divisors_antidiagonal_zero : divisors_antidiagonal 0 = ∅ := by { ext, si lemma divisors_antidiagonal_one : divisors_antidiagonal 1 = {(1,1)} := by { ext, simp [nat.mul_eq_one_iff, prod.ext_iff], } -lemma swap_mem_divisors_antidiagonal {x : ℕ × ℕ} (h : x ∈ divisors_antidiagonal n) : - x.swap ∈ divisors_antidiagonal n := -begin - rw [mem_divisors_antidiagonal, mul_comm] at h, - simp [h.1, h.2], -end +@[simp] lemma swap_mem_divisors_antidiagonal {x : ℕ × ℕ} : + x.swap ∈ divisors_antidiagonal n ↔ x ∈ divisors_antidiagonal n := +by rw [mem_divisors_antidiagonal, mem_divisors_antidiagonal, mul_comm, prod.swap] lemma fst_mem_divisors_of_mem_antidiagonal {x : ℕ × ℕ} (h : x ∈ divisors_antidiagonal n) : x.fst ∈ divisors n := @@ -204,19 +201,12 @@ end @[simp] lemma map_swap_divisors_antidiagonal : - (divisors_antidiagonal n).map ⟨prod.swap, prod.swap_right_inverse.injective⟩ - = divisors_antidiagonal n := + (divisors_antidiagonal n).map (equiv.prod_comm _ _).to_embedding = divisors_antidiagonal n := begin + rw [← coe_inj, coe_map, equiv.coe_to_embedding, equiv.coe_prod_comm, + set.image_swap_eq_preimage_swap], ext, - simp only [exists_prop, mem_divisors_antidiagonal, finset.mem_map, function.embedding.coe_fn_mk, - ne.def, prod.swap_prod_mk, prod.exists], - split, - { rintros ⟨x, y, ⟨⟨rfl, h⟩, rfl⟩⟩, - simp [mul_comm, h], }, - { rintros ⟨rfl, h⟩, - use [a.snd, a.fst], - rw mul_comm, - simp [h] } + exact swap_mem_divisors_antidiagonal, end @[simp] lemma image_fst_divisors_antidiagonal : @@ -233,26 +223,21 @@ end lemma map_div_right_divisors : n.divisors.map ⟨λ d, (d, n/d), λ p₁ p₂, congr_arg prod.fst⟩ = n.divisors_antidiagonal := begin - obtain rfl | hn := decidable.eq_or_ne n 0, - { simp }, ext ⟨d, nd⟩, - simp only [and_true, finset.mem_map, exists_eq_left, ne.def, hn, not_false_iff, - mem_divisors_antidiagonal, function.embedding.coe_fn_mk, mem_divisors], + simp only [mem_map, mem_divisors_antidiagonal, function.embedding.coe_fn_mk, mem_divisors, + prod.ext_iff, exists_prop, and.left_comm, exists_eq_left], split, - { rintro ⟨a, ⟨k, rfl⟩, h⟩, - obtain ⟨rfl, rfl⟩ := prod.mk.inj_iff.1 h, - have := (mul_ne_zero_iff.1 hn).1, - rw nat.mul_div_cancel_left _ (zero_lt_iff.mpr this), }, - { rintro rfl, - refine ⟨d, dvd_mul_right _ _, _⟩, - have := (mul_ne_zero_iff.1 hn).1, - rw nat.mul_div_cancel_left _ (zero_lt_iff.mpr this) } + { rintro ⟨⟨⟨k, rfl⟩, hn⟩, rfl⟩, + rw [nat.mul_div_cancel_left _ (left_ne_zero_of_mul hn).bot_lt], + exact ⟨rfl, hn⟩ }, + { rintro ⟨rfl, hn⟩, + exact ⟨⟨dvd_mul_right _ _, hn⟩, nat.mul_div_cancel_left _ (left_ne_zero_of_mul hn).bot_lt⟩ } end lemma map_div_left_divisors : n.divisors.map ⟨λ d, (n/d, d), λ p₁ p₂, congr_arg prod.snd⟩ = n.divisors_antidiagonal := begin - apply finset.map_injective ⟨prod.swap, prod.swap_right_inverse.injective⟩, + apply finset.map_injective (equiv.prod_comm _ _).to_embedding, rw [map_swap_divisors_antidiagonal, ←map_div_right_divisors, finset.map_map], refl, end From 758a6a4ee1fb83944f92c418ab32a62255002b43 Mon Sep 17 00:00:00 2001 From: Michail Karatarakis Date: Sat, 17 Dec 2022 16:15:00 +0000 Subject: [PATCH 0068/1291] feat(ring_theory/ideal/local_ring): add lemmas about local_ring.residue_field.map (#17969) Add `map_comp_residue` & `map_residue` in `residue_field.map`. Co-authored-by: Eric Wieser --- src/ring_theory/ideal/local_ring.lean | 6 ++++++ 1 file changed, 6 insertions(+) diff --git a/src/ring_theory/ideal/local_ring.lean b/src/ring_theory/ideal/local_ring.lean index 8928de69e1723..3f53ef4a90068 100644 --- a/src/ring_theory/ideal/local_ring.lean +++ b/src/ring_theory/ideal/local_ring.lean @@ -349,6 +349,12 @@ lemma map_comp (f : T →+* R) (g : R →+* S) [is_local_ring_hom f] [is_local_r (local_ring.residue_field.map g).comp (local_ring.residue_field.map f) := ideal.quotient.ring_hom_ext $ ring_hom.ext $ λx, rfl +lemma map_comp_residue (f : R →+* S) [is_local_ring_hom f] : + (residue_field.map f).comp (residue R) = (residue S).comp f := rfl + +lemma map_residue (f : R →+* S) [is_local_ring_hom f] (r : R) : + residue_field.map f (residue R r) = residue S (f r) := rfl + lemma map_id_apply (x : residue_field R) : map (ring_hom.id R) x = x := fun_like.congr_fun map_id x From dcf2250875895376a142faeeac5eabff32c48655 Mon Sep 17 00:00:00 2001 From: Anatole Dedecker Date: Sat, 17 Dec 2022 19:37:31 +0000 Subject: [PATCH 0069/1291] Split `infinite_of_mem_nhds` to extract a useful lemma (#17971) Nothing new mathematically since the new lemma is just the contraposition of `infinite_of_mem_nhds`, but the proof of the latter actually goes through a proof of the former so it makes more sense to reorganize things like that. Requested by @xroblot on [Zulip](https://leanprover.zulipchat.com/#narrow/stream/217875-Is-there-code-for-X.3F/topic/is_open.20.7Ba.7D.20from.20finite.20balls) --- src/topology/basic.lean | 5 +++++ src/topology/separation.lean | 20 ++++++++++++++------ 2 files changed, 19 insertions(+), 6 deletions(-) diff --git a/src/topology/basic.lean b/src/topology/basic.lean index 0d5ead6ede2da..ffcb7209afc38 100644 --- a/src/topology/basic.lean +++ b/src/topology/basic.lean @@ -995,6 +995,11 @@ begin simp [is_open_iff_nhds, h] } end +lemma is_open_singleton_iff_punctured_nhds {α : Type*} [topological_space α] (a : α) : + is_open ({a} : set α) ↔ (𝓝[≠] a) = ⊥ := +by rw [is_open_singleton_iff_nhds_eq_pure, nhds_within, ← mem_iff_inf_principal_compl, + ← le_pure_iff, nhds_ne_bot.le_pure_iff] + lemma mem_closure_iff_frequently {s : set α} {a : α} : a ∈ closure s ↔ ∃ᶠ x in 𝓝 a, x ∈ s := by rw [filter.frequently, filter.eventually, ← mem_interior_iff_mem_nhds, closure_eq_compl_interior_compl]; refl diff --git a/src/topology/separation.lean b/src/topology/separation.lean index 96c9ff01d93b7..3e55ac79f6534 100644 --- a/src/topology/separation.lean +++ b/src/topology/separation.lean @@ -663,17 +663,25 @@ show tendsto f (𝓝 a) (𝓝 $ f a), by rwa eq_of_tendsto_nhds h tendsto (λ x, c) l (𝓝 d) ↔ c = d := by simp_rw [tendsto, filter.map_const, pure_le_nhds_iff] -/-- If the punctured neighborhoods of a point form a nontrivial filter, then any neighborhood is -infinite. -/ -lemma infinite_of_mem_nhds {α} [topological_space α] [t1_space α] (x : α) [hx : ne_bot (𝓝[≠] x)] - {s : set α} (hs : s ∈ 𝓝 x) : set.infinite s := +/-- A point with a finite neighborhood has to be isolated. -/ +lemma is_open_singleton_of_finite_mem_nhds {α : Type*} [topological_space α] [t1_space α] + (x : α) {s : set α} (hs : s ∈ 𝓝 x) (hsf : s.finite) : is_open ({x} : set α) := begin - intro hsf, have A : {x} ⊆ s, by simp only [singleton_subset_iff, mem_of_mem_nhds hs], have B : is_closed (s \ {x}) := (hsf.subset (diff_subset _ _)).is_closed, have C : (s \ {x})ᶜ ∈ 𝓝 x, from B.is_open_compl.mem_nhds (λ h, h.2 rfl), have D : {x} ∈ 𝓝 x, by simpa only [← diff_eq, diff_diff_cancel_left A] using inter_mem hs C, - rwa [← mem_interior_iff_mem_nhds, interior_singleton] at D + rwa [← mem_interior_iff_mem_nhds, ← singleton_subset_iff, subset_interior_iff_is_open] at D +end + +/-- If the punctured neighborhoods of a point form a nontrivial filter, then any neighborhood is +infinite. -/ +lemma infinite_of_mem_nhds {α} [topological_space α] [t1_space α] (x : α) [hx : ne_bot (𝓝[≠] x)] + {s : set α} (hs : s ∈ 𝓝 x) : set.infinite s := +begin + refine λ hsf, hx.1 _, + rw [← is_open_singleton_iff_punctured_nhds], + exact is_open_singleton_of_finite_mem_nhds x hs hsf end lemma discrete_of_t1_of_finite {X : Type*} [topological_space X] [t1_space X] [finite X] : From c5c7e2760814660967bc27f0de95d190a22297f3 Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Sat, 17 Dec 2022 21:41:32 +0000 Subject: [PATCH 0070/1291] feat(data/polynomial/basic): add C_ne_zero (#17970) --- src/data/polynomial/basic.lean | 2 ++ 1 file changed, 2 insertions(+) diff --git a/src/data/polynomial/basic.lean b/src/data/polynomial/basic.lean index 75af008301507..8b825846cbcb2 100644 --- a/src/data/polynomial/basic.lean +++ b/src/data/polynomial/basic.lean @@ -497,6 +497,8 @@ lemma C_injective : injective (C : R → R[X]) := monomial_injective 0 @[simp] lemma C_inj : C a = C b ↔ a = b := C_injective.eq_iff @[simp] lemma C_eq_zero : C a = 0 ↔ a = 0 := C_injective.eq_iff' (map_zero C) +lemma C_ne_zero : C a ≠ 0 ↔ a ≠ 0 := C_eq_zero.not + lemma subsingleton_iff_subsingleton : subsingleton R[X] ↔ subsingleton R := ⟨@injective.subsingleton _ _ _ C_injective, by { introI, apply_instance } ⟩ From bfc28c4bd321337c094b2fc9353ea4ad9f3b06a4 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sun, 18 Dec 2022 09:36:25 +0000 Subject: [PATCH 0071/1291] feat(measure_theory/group/fundamental_domain): Interior of a fundamental domain (#17533) Define the interior and frontier of a fundamental domain. Co-authored-by: Alex- J. Best --- src/measure_theory/group/action.lean | 2 + .../group/fundamental_domain.lean | 145 ++++++++++++++++++ 2 files changed, 147 insertions(+) diff --git a/src/measure_theory/group/action.lean b/src/measure_theory/group/action.lean index a0c718bdc1c63..250c467310882 100644 --- a/src/measure_theory/group/action.lean +++ b/src/measure_theory/group/action.lean @@ -148,6 +148,8 @@ variable {μ} by simpa only [← preimage_smul_inv] using hs.preimage (measure_preserving_smul _ _).quasi_measure_preserving +lemma measure_smul_null {s} (h : μ s = 0) (c : G) : μ (c • s) = 0 := by rwa measure_smul + section is_minimal variables (G) [topological_space α] [has_continuous_const_smul G α] [mul_action.is_minimal G α] diff --git a/src/measure_theory/group/fundamental_domain.lean b/src/measure_theory/group/fundamental_domain.lean index 75d690a89c0ba..4d48c958b734d 100644 --- a/src/measure_theory/group/fundamental_domain.lean +++ b/src/measure_theory/group/fundamental_domain.lean @@ -24,6 +24,15 @@ fundamental domains have the same measure, and for a `G`-invariant function, its two fundamental domains are equal to each other. We also generate additive versions of all theorems in this file using the `to_additive` attribute. + +## Main declarations + +* `measure_theory.is_fundamental_domain`: Predicate for a set to be a fundamental domain of the + action of a group +* `measure_theory.fundamental_frontier`: Fundamental frontier of a set under the action of a group. + Elements of `s` that belong to some other translate of `s`. +* `measure_theory.fundamental_interior`: Fundamental interior of a set under the action of a group. + Elements of `s` that do not belong to any other translate of `s`. -/ open_locale ennreal pointwise topological_space nnreal ennreal measure_theory @@ -405,4 +414,140 @@ end end is_fundamental_domain +/-! ### Interior/frontier of a fundamental domain -/ + +section measurable_space +variables (G) [group G] [mul_action G α] (s : set α) {x : α} + +/-- The boundary of a fundamental domain, those points of the domain that also lie in a nontrivial +translate. -/ +@[to_additive measure_theory.add_fundamental_frontier "The boundary of a fundamental domain, those +points of the domain that also lie in a nontrivial translate."] +def fundamental_frontier : set α := s ∩ ⋃ (g : G) (hg : g ≠ 1), g • s + +/-- The interior of a fundamental domain, those points of the domain not lying in any translate. -/ +@[to_additive measure_theory.add_fundamental_interior "The interior of a fundamental domain, those +points of the domain not lying in any translate."] +def fundamental_interior : set α := s \ ⋃ (g : G) (hg : g ≠ 1), g • s + +variables {G s} + +@[simp, to_additive measure_theory.mem_add_fundamental_frontier] +lemma mem_fundamental_frontier : + x ∈ fundamental_frontier G s ↔ x ∈ s ∧ ∃ (g : G) (hg : g ≠ 1), x ∈ g • s := +by simp [fundamental_frontier] + +@[simp, to_additive measure_theory.mem_add_fundamental_interior] +lemma mem_fundamental_interior : + x ∈ fundamental_interior G s ↔ x ∈ s ∧ ∀ (g : G) (hg : g ≠ 1), x ∉ g • s := +by simp [fundamental_interior] + +@[to_additive measure_theory.add_fundamental_frontier_subset] +lemma fundamental_frontier_subset : fundamental_frontier G s ⊆ s := inter_subset_left _ _ + +@[to_additive measure_theory.add_fundamental_interior_subset] +lemma fundamental_interior_subset : fundamental_interior G s ⊆ s := diff_subset _ _ + +variables (G s) + +@[to_additive measure_theory.disjoint_add_fundamental_interior_add_fundamental_frontier] +lemma disjoint_fundamental_interior_fundamental_frontier : + disjoint (fundamental_interior G s) (fundamental_frontier G s) := +disjoint_sdiff_self_left.mono_right inf_le_right + +@[simp, to_additive measure_theory.add_fundamental_interior_union_add_fundamental_frontier] +lemma fundamental_interior_union_fundamental_frontier : + fundamental_interior G s ∪ fundamental_frontier G s = s := +diff_union_inter _ _ + +@[simp, to_additive measure_theory.add_fundamental_interior_union_add_fundamental_frontier] +lemma fundamental_frontier_union_fundamental_interior : + fundamental_frontier G s ∪ fundamental_interior G s = s := +inter_union_diff _ _ + +@[simp, to_additive measure_theory.sdiff_add_fundamental_interior] +lemma sdiff_fundamental_interior : s \ fundamental_interior G s = fundamental_frontier G s := +sdiff_sdiff_right_self + +@[simp, to_additive measure_theory.sdiff_add_fundamental_frontier] +lemma sdiff_fundamental_frontier : s \ fundamental_frontier G s = fundamental_interior G s := +diff_self_inter + +@[simp, to_additive measure_theory.add_fundamental_frontier_vadd] +lemma fundamental_frontier_smul [group H] [mul_action H α] [smul_comm_class H G α] (g : H) : + fundamental_frontier G (g • s) = g • fundamental_frontier G s := +by simp_rw [fundamental_frontier, smul_set_inter, smul_set_Union, smul_comm g] + +@[simp, to_additive measure_theory.add_fundamental_interior_vadd] +lemma fundamental_interior_smul [group H] [mul_action H α] [smul_comm_class H G α] (g : H) : + fundamental_interior G (g • s) = g • fundamental_interior G s := +by simp_rw [fundamental_interior, smul_set_sdiff, smul_set_Union, smul_comm g] + +@[to_additive measure_theory.pairwise_disjoint_add_fundamental_interior] +lemma pairwise_disjoint_fundamental_interior : + pairwise (disjoint on λ g : G, g • fundamental_interior G s) := +begin + refine λ a b hab, disjoint_left.2 _, + rintro _ ⟨x, hx, rfl⟩ ⟨y, hy, hxy⟩, + rw mem_fundamental_interior at hx hy, + refine hx.2 (a⁻¹ * b) _ _, + rwa [ne.def, inv_mul_eq_iff_eq_mul, mul_one, eq_comm], + simpa [mul_smul, ←hxy, mem_inv_smul_set_iff] using hy.1, +end + +variables [countable G] [measurable_space G] [measurable_space α] [has_measurable_smul G α] + {μ : measure α} [smul_invariant_measure G α μ] + +@[to_additive measure_theory.null_measurable_set.add_fundamental_frontier] +protected lemma null_measurable_set.fundamental_frontier (hs : null_measurable_set s μ) : + null_measurable_set (fundamental_frontier G s) μ := +hs.inter $ null_measurable_set.Union $ λ g, null_measurable_set.Union $ λ hg, hs.smul _ + +@[to_additive measure_theory.null_measurable_set.add_fundamental_interior] +protected lemma null_measurable_set.fundamental_interior (hs : null_measurable_set s μ) : + null_measurable_set (fundamental_interior G s) μ := +hs.diff $ null_measurable_set.Union $ λ g, null_measurable_set.Union $ λ hg, hs.smul _ + +end measurable_space + +namespace is_fundamental_domain +section group +variables [countable G] [group G] [mul_action G α] [measurable_space α] {μ : measure α} {s : set α} + (hs : is_fundamental_domain G s μ) +include hs + +@[to_additive measure_theory.is_add_fundamental_domain.measure_add_fundamental_frontier] +lemma measure_fundamental_frontier : μ (fundamental_frontier G s) = 0 := +by simpa only [fundamental_frontier, Union₂_inter, measure_Union_null_iff', one_smul, + measure_Union_null_iff, inter_comm s, function.on_fun] using λ g (hg : g ≠ 1), hs.ae_disjoint hg + +@[to_additive measure_theory.is_add_fundamental_domain.measure_add_fundamental_interior] +lemma measure_fundamental_interior : μ (fundamental_interior G s) = μ s := +measure_diff_null' hs.measure_fundamental_frontier + +end group + +variables [countable G] [group G] [mul_action G α] [measurable_space α] {μ : measure α} {s : set α} + (hs : is_fundamental_domain G s μ) [measurable_space G] [has_measurable_smul G α] + [smul_invariant_measure G α μ] +include hs + +protected lemma fundamental_interior : is_fundamental_domain G (fundamental_interior G s) μ := +{ null_measurable_set := hs.null_measurable_set.fundamental_interior _ _, + ae_covers := begin + simp_rw [ae_iff, not_exists, ←mem_inv_smul_set_iff, set_of_forall, ←compl_set_of, set_of_mem_eq, + ←compl_Union], + have : (⋃ g : G, g⁻¹ • s) \ (⋃ g : G, g⁻¹ • fundamental_frontier G s) ⊆ + ⋃ g : G, g⁻¹ • fundamental_interior G s, + { simp_rw [diff_subset_iff, ←Union_union_distrib, ←smul_set_union, + fundamental_frontier_union_fundamental_interior] }, + refine eq_bot_mono (μ.mono $ compl_subset_compl.2 this) _, + simp only [Union_inv_smul, outer_measure.measure_of_eq_coe, coe_to_outer_measure, compl_sdiff, + ennreal.bot_eq_zero, himp_eq, sup_eq_union, @Union_smul_eq_set_of_exists _ _ _ _ s], + exact measure_union_null + (measure_Union_null $ λ _, measure_smul_null hs.measure_fundamental_frontier _) hs.ae_covers, + end, + ae_disjoint := (pairwise_disjoint_fundamental_interior _ _).mono $ λ _ _, disjoint.ae_disjoint } + +end is_fundamental_domain end measure_theory From 04206afe885f3bbf96ea5fcee5ee04fcf32fdbab Mon Sep 17 00:00:00 2001 From: loefflerd Date: Sun, 18 Dec 2022 11:38:17 +0000 Subject: [PATCH 0072/1291] feat(topology/instances): add_circle = interval with ends identified (#17889) --- src/topology/instances/add_circle.lean | 184 +++++++++++++++++++++---- 1 file changed, 159 insertions(+), 25 deletions(-) diff --git a/src/topology/instances/add_circle.lean b/src/topology/instances/add_circle.lean index ddec749e13e99..a7f8b9a387576 100644 --- a/src/topology/instances/add_circle.lean +++ b/src/topology/instances/add_circle.lean @@ -19,14 +19,19 @@ We define the additive circle `add_circle p` as the quotient `𝕜 ⧸ (ℤ ∙ See also `circle` and `real.angle`. For the normed group structure on `add_circle`, see `add_circle.normed_add_comm_group` in a later file. -## Main definitions: +## Main definitions and results: * `add_circle`: the additive circle `𝕜 ⧸ (ℤ ∙ p)` for some period `p : 𝕜` * `unit_add_circle`: the special case `ℝ ⧸ ℤ` * `add_circle.equiv_add_circle`: the rescaling equivalence `add_circle p ≃+ add_circle q` - * `add_circle.equiv_Ico`: the natural equivalence `add_circle p ≃ Ico 0 p` + * `add_circle.equiv_Ico`: the natural equivalence `add_circle p ≃ Ico a (a + p)` * `add_circle.add_order_of_div_of_gcd_eq_one`: rational points have finite order * `add_circle.exists_gcd_eq_one_of_is_of_fin_add_order`: finite-order points are rational + * `add_circle.homeo_Icc_quot`: the natural topological equivalence between `add_circle p` and + `Icc a (a + p)` with its endpoints identified. + * `add_circle.lift_Ico_continuous`: if `f : ℝ → B` is continuous, and `f a = f (a + p)` for + some `a`, then there is a continuous function `add_circle p → B` which agrees with `f` on + `Icc a (a + p)`. ## Implementation notes: @@ -43,9 +48,9 @@ the rational circle `add_circle (1 : ℚ)`, and so we set things up more general noncomputable theory -open set add_subgroup topological_space +open set function add_subgroup topological_space -variables {𝕜 : Type*} +variables {𝕜 : Type*} {B : Type*} /-- The "additive circle": `𝕜 ⧸ (ℤ ∙ p)`. See also `circle` and `real.angle`. -/ @[derive [add_comm_group, topological_space, topological_add_group, inhabited, has_coe_t 𝕜], @@ -80,6 +85,12 @@ begin { exact ⟨(n : ℤ), by simp⟩, }, end +@[simp] lemma coe_add_period (x : 𝕜) : (((x + p) : 𝕜) : add_circle p) = x := +begin + rw [quotient_add_group.coe_add, ←eq_sub_iff_add_eq', sub_self, quotient_add_group.eq_zero_iff], + exact mem_zmultiples p, +end + @[continuity, nolint unused_arguments] protected lemma continuous_mk' : continuous (quotient_add_group.mk' (zmultiples p) : 𝕜 → add_circle p) := continuous_coinduced_rng @@ -87,24 +98,49 @@ continuous_coinduced_rng variables [hp : fact (0 < p)] include hp -variables [archimedean 𝕜] +variables (a : 𝕜) [archimedean 𝕜] + +/-- The natural equivalence between `add_circle p` and the half-open interval `[a, a + p)`. -/ +def equiv_Ico : add_circle p ≃ Ico a (a + p) := quotient_add_group.equiv_Ico_mod a hp.out + +/-- Given a function on `[a, a + p)`, lift it to `add_circle p`. -/ +def lift_Ico (f : 𝕜 → B) : add_circle p → B := restrict _ f ∘ add_circle.equiv_Ico p a + +variables {p a} + +lemma coe_eq_coe_iff_of_mem_Ico {x y : 𝕜} + (hx : x ∈ Ico a (a + p)) (hy : y ∈ Ico a (a + p)) : (x : add_circle p) = y ↔ x = y := +begin + refine ⟨λ h, _, by tauto⟩, + suffices : (⟨x, hx⟩ : Ico a (a + p)) = ⟨y, hy⟩, by exact subtype.mk.inj this, + apply_fun equiv_Ico p a at h, + rw [←(equiv_Ico p a).right_inv ⟨x, hx⟩, ←(equiv_Ico p a).right_inv ⟨y, hy⟩], + exact h +end + +lemma lift_Ico_coe_apply {f : 𝕜 → B} {x : 𝕜} (hx : x ∈ Ico a (a + p)) : lift_Ico p a f ↑x = f x := +begin + have : (equiv_Ico p a) x = ⟨x, hx⟩, + { rw equiv.apply_eq_iff_eq_symm_apply, + refl, }, + rw [lift_Ico, comp_apply, this], + refl, +end -/-- The natural equivalence between `add_circle p` and the half-open interval `[0, p)`. -/ -def equiv_Ico : add_circle p ≃ Ico 0 p := -(quotient_add_group.equiv_Ico_mod 0 hp.out).trans $ equiv.set.of_eq $ by rw zero_add +variables (p a) -@[continuity] lemma continuous_equiv_Ico_symm : continuous (equiv_Ico p).symm := +@[continuity] lemma continuous_equiv_Ico_symm : continuous (equiv_Ico p a).symm := continuous_quotient_mk.comp continuous_subtype_coe -/-- The image of the closed-open interval `[0, p)` under the quotient map `𝕜 → add_circle p` is the -entire space. -/ -@[simp] lemma coe_image_Ico_eq : (coe : 𝕜 → add_circle p) '' Ico 0 p = univ := -by { rw image_eq_range, exact (equiv_Ico p).symm.range_eq_univ } +/-- The image of the closed-open interval `[a, a + p)` under the quotient map `𝕜 → add_circle p` is +the entire space. -/ +@[simp] lemma coe_image_Ico_eq : (coe : 𝕜 → add_circle p) '' Ico a (a + p) = univ := +by { rw image_eq_range, exact (equiv_Ico p a).symm.range_eq_univ } /-- The image of the closed interval `[0, p]` under the quotient map `𝕜 → add_circle p` is the entire space. -/ -@[simp] lemma coe_image_Icc_eq : (coe : 𝕜 → add_circle p) '' Icc 0 p = univ := -eq_top_mono (image_subset _ Ico_subset_Icc_self) $ coe_image_Ico_eq _ +@[simp] lemma coe_image_Icc_eq : (coe : 𝕜 → add_circle p) '' Icc a (a + p) = univ := +eq_top_mono (image_subset _ Ico_subset_Icc_self) $ coe_image_Ico_eq _ _ end linear_ordered_add_comm_group @@ -133,19 +169,19 @@ section floor_ring variables [floor_ring 𝕜] @[simp] lemma coe_equiv_Ico_mk_apply (x : 𝕜) : - (equiv_Ico p $ quotient_add_group.mk x : 𝕜) = int.fract (x / p) * p := + (equiv_Ico p 0 $ quotient_add_group.mk x : 𝕜) = int.fract (x / p) * p := to_Ico_mod_eq_fract_mul _ x instance : divisible_by (add_circle p) ℤ := -{ div := λ x n, (↑(((n : 𝕜)⁻¹) * (equiv_Ico p x : 𝕜)) : add_circle p), +{ div := λ x n, (↑(((n : 𝕜)⁻¹) * (equiv_Ico p 0 x : 𝕜)) : add_circle p), div_zero := λ x, by simp only [algebra_map.coe_zero, quotient_add_group.coe_zero, inv_zero, zero_mul], div_cancel := λ n x hn, begin replace hn : (n : 𝕜) ≠ 0, { norm_cast, assumption, }, - change n • quotient_add_group.mk' _ ((n : 𝕜)⁻¹ * ↑(equiv_Ico p x)) = x, + change n • quotient_add_group.mk' _ ((n : 𝕜)⁻¹ * ↑(equiv_Ico p 0 x)) = x, rw [← map_zsmul, ← smul_mul_assoc, zsmul_eq_mul, mul_inv_cancel hn, one_mul], - exact (equiv_Ico p).symm_apply_apply x, + exact (equiv_Ico p 0).symm_apply_apply x, end, } end floor_ring @@ -224,11 +260,11 @@ begin change ∃ m, gcd m n = 1 ∧ m < n ∧ ↑((↑m / ↑n) * p) = u, have hn : 0 < n := add_order_of_pos' h, have hn₀ : (n : 𝕜) ≠ 0, { norm_cast, exact ne_of_gt hn, }, - let x := (equiv_Ico p u : 𝕜), - have hxu : (x : add_circle p) = u := (equiv_Ico p).symm_apply_apply u, + let x := (equiv_Ico p 0 u : 𝕜), + have hxu : (x : add_circle p) = u := (equiv_Ico p 0).symm_apply_apply u, have hx₀ : 0 < (add_order_of (x : add_circle p)), { rw ← hxu at h, exact add_order_of_pos' h, }, have hx₁ : 0 < x, - { refine lt_of_le_of_ne (equiv_Ico p u).2.1 _, + { refine lt_of_le_of_ne (equiv_Ico p 0 u).2.1 _, contrapose! hu, rw [← hxu, ← hu, quotient_add_group.coe_zero], }, obtain ⟨m, hm : m • p = add_order_of ↑x • x⟩ := (coe_eq_zero_of_pos_iff p hp.out @@ -239,8 +275,9 @@ begin refine ⟨m, (_ : gcd m n = 1), (_ : m < n), hux⟩, { have := gcd_mul_add_order_of_div_eq p m hn, rwa [hux, nat.mul_left_eq_self_iff hn] at this, }, - { have : n • x < n • p := smul_lt_smul_of_pos (equiv_Ico p u).2.2 hn, - rwa [nsmul_eq_mul, nsmul_eq_mul, ← hm, mul_lt_mul_right hp.out, nat.cast_lt] at this, }, + { have : n • x < n • p := smul_lt_smul_of_pos _ hn, + rwa [nsmul_eq_mul, nsmul_eq_mul, ← hm, mul_lt_mul_right hp.out, nat.cast_lt] at this, + simpa [zero_add] using (equiv_Ico p 0 u).2.2, }, end end finite_order_points @@ -252,7 +289,7 @@ variables (p : ℝ) /-- The "additive circle" `ℝ ⧸ (ℤ ∙ p)` is compact. -/ instance compact_space [fact (0 < p)] : compact_space $ add_circle p := begin - rw [← is_compact_univ_iff, ← coe_image_Icc_eq p], + rw [← is_compact_univ_iff, ← coe_image_Icc_eq p 0], exact is_compact_Icc.image (add_circle.continuous_mk' p), end @@ -279,3 +316,100 @@ local attribute [instance] fact_zero_lt_one /-- The unit circle `ℝ ⧸ ℤ`. -/ @[derive [compact_space, normal_space, second_countable_topology]] abbreviation unit_add_circle := add_circle (1 : ℝ) + +section identify_Icc_ends +/-! This section proves that for any `a`, the natural map from `[a, a + p] ⊂ ℝ` to `add_circle p` +gives an identification of `add_circle p`, as a topological space, with the quotient of `[a, a + p]` +by the equivalence relation identifying the endpoints. -/ + +namespace add_circle + +section linear_ordered_add_comm_group + +variables [linear_ordered_add_comm_group 𝕜] [topological_space 𝕜] [order_topology 𝕜] +(p a : 𝕜) [hp : fact (0 < p)] + +include hp + +local notation `𝕋` := add_circle p + +/-- The relation identifying the endpoints of `Icc a (a + p)`. -/ +inductive endpoint_ident : Icc a (a + p) → Icc a (a + p) → Prop +| mk : endpoint_ident + ⟨a, left_mem_Icc.mpr $ le_add_of_nonneg_right hp.out.le⟩ + ⟨a + p, right_mem_Icc.mpr $ le_add_of_nonneg_right hp.out.le⟩ + +variables [archimedean 𝕜] + +/-- The equivalence between `add_circle p` and the quotient of `[a, a + p]` by the relation +identifying the endpoints. -/ +def equiv_Icc_quot : 𝕋 ≃ quot (endpoint_ident p a) := +{ to_fun := λ x, quot.mk _ $ subtype.map id Ico_subset_Icc_self (equiv_Ico _ _ x), + inv_fun := λ x, quot.lift_on x coe $ by { rintro _ _ ⟨_⟩, exact (coe_add_period p a).symm }, + left_inv := (equiv_Ico p a).symm_apply_apply, + right_inv := quot.ind $ by + { rintro ⟨x, hx⟩, + have := _, + rcases ne_or_eq x (a + p) with h | rfl, + { revert x, exact this }, + { rw ← quot.sound endpoint_ident.mk, exact this _ _ (lt_add_of_pos_right a hp.out).ne }, + intros x hx h, + congr, ext1, + apply congr_arg subtype.val ((equiv_Ico p a).right_inv ⟨x, hx.1, hx.2.lt_of_ne h⟩) } } + +end linear_ordered_add_comm_group + +section real + +variables {p a : ℝ} [hp : fact (0 < p)] +include hp + +local notation `𝕋` := add_circle p + +/- doesn't work if inlined in `homeo_of_equiv_compact_to_t2` -- why? -/ +private lemma continuous_equiv_Icc_quot_symm : continuous (equiv_Icc_quot p a).symm := +continuous_quot_lift _ $ (add_circle.continuous_mk' p).comp continuous_subtype_coe + +/-- The natural map from `[a, a + p] ⊂ ℝ` with endpoints identified to `ℝ / ℤ • p`, as a +homeomorphism of topological spaces. -/ +def homeo_Icc_quot : 𝕋 ≃ₜ quot (endpoint_ident p a) := +(continuous.homeo_of_equiv_compact_to_t2 continuous_equiv_Icc_quot_symm).symm + +/-! We now show that a continuous function on `[a, a + p]` satisfying `f a = f (a + p)` is +the pullback of a continuous function on `unit_add_circle`. -/ + +lemma eq_of_end_ident {f : ℝ → B} (hf : f a = f (a + p)) (x y : Icc a (a + p)) : + endpoint_ident p a x y → f x = f y := by { rintro ⟨_⟩, exact hf } + +lemma lift_Ico_eq_lift_Icc {f : ℝ → B} (h : f a = f (a + p)) : + lift_Ico p a f = (quot.lift (restrict (Icc a $ a + p) f) $ eq_of_end_ident h) + ∘ equiv_Icc_quot p a := +funext (λ x, by refl) + +lemma lift_Ico_continuous [topological_space B] {f : ℝ → B} (hf : f a = f (a + p)) + (hc : continuous_on f $ Icc a (a + p)) : continuous (lift_Ico p a f) := +begin + rw lift_Ico_eq_lift_Icc hf, + refine continuous.comp _ homeo_Icc_quot.continuous_to_fun, + exact continuous_coinduced_dom.mpr (continuous_on_iff_continuous_restrict.mp hc), +end + +end real + +section zero_based + +variables {p : ℝ} [hp : fact (0 < p)] +include hp + +lemma lift_Ico_zero_coe_apply {f : ℝ → B} {x : ℝ} (hx : x ∈ Ico 0 p) : + lift_Ico p 0 f ↑x = f x := lift_Ico_coe_apply (by rwa zero_add) + +lemma lift_Ico_zero_continuous [topological_space B] {f : ℝ → B} + (hf : f 0 = f p) (hc : continuous_on f $ Icc 0 p) : continuous (lift_Ico p 0 f) := +lift_Ico_continuous (by rwa zero_add : f 0 = f (0 + p)) (by rwa zero_add) + +end zero_based + +end add_circle + +end identify_Icc_ends From 138f1db8dbd28f2a49c8dc0a67d568a71573869b Mon Sep 17 00:00:00 2001 From: Felix-Weilacher Date: Sun, 18 Dec 2022 16:44:53 +0000 Subject: [PATCH 0073/1291] feat(topology): add basics of perfect sets (#17492) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit add a file perfect.lean which contains basic definitions and facts about perfect sets as well as a version of the Cantor-Bendixson theorem. This is intended to build up to a proof of the Borel isomorphism theorem. Co-authored-by: Rémy Degenne Co-authored-by: Felix-Weilacher <112423742+Felix-Weilacher@users.noreply.github.com> Co-authored-by: RemyDegenne --- src/topology/basic.lean | 35 +++++++ src/topology/perfect.lean | 213 ++++++++++++++++++++++++++++++++++++++ 2 files changed, 248 insertions(+) create mode 100644 src/topology/perfect.lean diff --git a/src/topology/basic.lean b/src/topology/basic.lean index ffcb7209afc38..101947cac8602 100644 --- a/src/topology/basic.lean +++ b/src/topology/basic.lean @@ -771,6 +771,14 @@ eventually_nhds_iff.2 ⟨t, λ x hx, eventually_nhds_iff.2 ⟨t, htp, hto, hx⟩ (∀ᶠ y in 𝓝 a, ∀ᶠ x in 𝓝 y, p x) ↔ ∀ᶠ x in 𝓝 a, p x := ⟨λ h, h.self_of_nhds, λ h, h.eventually_nhds⟩ +@[simp] lemma frequently_frequently_nhds {p : α → Prop} {a : α} : + (∃ᶠ y in 𝓝 a, ∃ᶠ x in 𝓝 y, p x) ↔ (∃ᶠ x in 𝓝 a, p x) := +begin + rw ← not_iff_not, + simp_rw not_frequently, + exact eventually_eventually_nhds, +end + @[simp] lemma eventually_mem_nhds {s : set α} {a : α} : (∀ᶠ x in 𝓝 a, s ∈ 𝓝 x) ↔ s ∈ 𝓝 a := eventually_eventually_nhds @@ -938,6 +946,33 @@ begin exact ne_bot_of_le this end +/--A point `x` is an accumulation point of a filter `F` if `𝓝[≠] x ⊓ F ≠ ⊥`.-/ +def acc_pt (x : α) (F : filter α) : Prop := ne_bot (𝓝[≠] x ⊓ F) + +lemma acc_iff_cluster (x : α) (F : filter α) : acc_pt x F ↔ cluster_pt x (𝓟 {x}ᶜ ⊓ F) := +by rw [acc_pt, nhds_within, cluster_pt, inf_assoc] + +/-- `x` is an accumulation point of a set `C` iff it is a cluster point of `C ∖ {x}`.-/ +lemma acc_principal_iff_cluster (x : α) (C : set α) : + acc_pt x (𝓟 C) ↔ cluster_pt x (𝓟(C \ {x})) := +by rw [acc_iff_cluster, inf_principal, inter_comm]; refl + +/-- `x` is an accumulation point of a set `C` iff every neighborhood +of `x` contains a point of `C` other than `x`. -/ +lemma acc_pt_iff_nhds (x : α) (C : set α) : acc_pt x (𝓟 C) ↔ ∀ U ∈ 𝓝 x, ∃ y ∈ U ∩ C, y ≠ x := +by simp [acc_principal_iff_cluster, cluster_pt_principal_iff, set.nonempty, exists_prop, + and_assoc, and_comm (¬ _ = x)] + +/-- `x` is an accumulation point of a set `C` iff +there are points near `x` in `C` and different from `x`.-/ +lemma acc_pt_iff_frequently (x : α) (C : set α) : acc_pt x (𝓟 C) ↔ ∃ᶠ y in 𝓝 x, y ≠ x ∧ y ∈ C := +by simp [acc_principal_iff_cluster, cluster_pt_principal_iff_frequently, and_comm] + +/-- If `x` is an accumulation point of `F` and `F ≤ G`, then +`x` is an accumulation point of `D. -/ +lemma acc_pt.mono {x : α} {F G : filter α} (h : acc_pt x F) (hFG : F ≤ G) : acc_pt x G := +⟨ne_bot_of_le_ne_bot h.ne (inf_le_inf_left _ hFG)⟩ + /-! ### Interior, closure and frontier in terms of neighborhoods -/ diff --git a/src/topology/perfect.lean b/src/topology/perfect.lean new file mode 100644 index 0000000000000..f788259dc54ae --- /dev/null +++ b/src/topology/perfect.lean @@ -0,0 +1,213 @@ +/- +Copyright (c) 2022 Felix Weilacher. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Felix Weilacher +-/ +import topology.separation +import topology.bases + +/-! +# Perfect Sets + +In this file we define perfect subsets of a topological space, and prove some basic properties, +including a version of the Cantor-Bendixson Theorem. + +## Main Definitions + +* `perfect C`: A set `C` is perfect, meaning it is closed and every point of it + is an accumulation point of itself. + +## Main Statements + +* `perfect.splitting`: A perfect nonempty set contains two disjoint perfect nonempty subsets. + The main inductive step in the construction of an embedding from the Cantor space to a + perfect nonempty complete metric space. +* `exists_countable_union_perfect_of_is_closed`: One version of the **Cantor-Bendixson Theorem**: + A closed set in a second countable space can be written as the union of a countable set and a + perfect set. + +## Implementation Notes + +We do not require perfect sets to be nonempty. + +We define a nonstandard predicate, `preperfect`, which drops the closed-ness requirement +from the definition of perfect. In T1 spaces, this is equivalent to having a perfect closure, +see `preperfect_iff_perfect_closure`. + +## References + +* [kechris1995] (Chapter 6) + +## Tags + +accumulation point, perfect set, Cantor-Bendixson. + +-/ + +open_locale topological_space filter +open topological_space filter set + +variables {α : Type*} [topological_space α] {C : set α} + +/-- If `x` is an accumulation point of a set `C` and `U` is a neighborhood of `x`, +then `x` is an accumulation point of `U ∩ C`. -/ +theorem acc_pt.nhds_inter {x : α} {U : set α} (h_acc : acc_pt x (𝓟 C)) (hU : U ∈ 𝓝 x) : + acc_pt x (𝓟 (U ∩ C)) := +begin + have : 𝓝[≠] x ≤ 𝓟 U, + { rw le_principal_iff, + exact mem_nhds_within_of_mem_nhds hU, }, + rw [acc_pt, ← inf_principal, ← inf_assoc, inf_of_le_left this], + exact h_acc, +end + +/-- A set `C` is preperfect if all of its points are accumulation points of itself. +If `C` is nonempty and `α` is a T1 space, this is equivalent to the closure of `C` being perfect. +See `preperfect_iff_perfect_closure`.-/ +def preperfect (C : set α) : Prop := ∀ x ∈ C, acc_pt x (𝓟 C) + +/-- A set `C` is called perfect if it is closed and all of its +points are accumulation points of itself. +Note that we do not require `C` to be nonempty.-/ +structure perfect (C : set α) : Prop := +(closed : is_closed C) +(acc : preperfect C) + +lemma preperfect_iff_nhds : preperfect C ↔ ∀ x ∈ C, ∀ U ∈ 𝓝 x, ∃ y ∈ U ∩ C, y ≠ x := +by simp only [preperfect, acc_pt_iff_nhds] + +/-- The intersection of a preperfect set and an open set is preperfect-/ +theorem preperfect.open_inter {U : set α} (hC : preperfect C) (hU : is_open U) : + preperfect (U ∩ C) := +begin + rintros x ⟨xU, xC⟩, + apply (hC _ xC).nhds_inter, + exact hU.mem_nhds xU, +end + +/-- The closure of a preperfect set is perfect. +For a converse, see `preperfect_iff_perfect_closure`-/ +theorem preperfect.perfect_closure (hC : preperfect C) : perfect (closure C) := +begin + split, { exact is_closed_closure }, + intros x hx, + by_cases h : x ∈ C; apply acc_pt.mono _ (principal_mono.mpr subset_closure), + { exact hC _ h }, + have : {x}ᶜ ∩ C = C := by simp [h], + rw [acc_pt, nhds_within, inf_assoc, inf_principal, this], + rw [closure_eq_cluster_pts] at hx, + exact hx, +end + +/-- In a T1 space, being preperfect is equivalent to having perfect closure.-/ +theorem preperfect_iff_perfect_closure [t1_space α] : + preperfect C ↔ perfect (closure C) := +begin + split; intro h, { exact h.perfect_closure }, + intros x xC, + have H : acc_pt x (𝓟 (closure C)) := h.acc _ (subset_closure xC), + rw acc_pt_iff_frequently at *, + have : ∀ y , y ≠ x ∧ y ∈ closure C → ∃ᶠ z in 𝓝 y, z ≠ x ∧ z ∈ C, + { rintros y ⟨hyx, yC⟩, + simp only [← mem_compl_singleton_iff, @and_comm _ (_ ∈ C) , ← frequently_nhds_within_iff, + hyx.nhds_within_compl_singleton, ← mem_closure_iff_frequently], + exact yC, }, + rw ← frequently_frequently_nhds, + exact H.mono this, +end + +theorem perfect.closure_nhds_inter {U : set α} (hC : perfect C) (x : α) (xC : x ∈ C) (xU : x ∈ U) + (Uop : is_open U) : perfect (closure (U ∩ C)) ∧ (closure (U ∩ C)).nonempty := +begin + split, + { apply preperfect.perfect_closure, + exact (hC.acc).open_inter Uop, }, + apply nonempty.closure, + exact ⟨x, ⟨xU, xC⟩⟩, +end + +/-- Given a perfect nonempty set in a T2.5 space, we can find two disjoint perfect subsets +This is the main inductive step in the proof of the Cantor-Bendixson Theorem-/ +lemma perfect.splitting [t2_5_space α] (hC : perfect C) (hnonempty : C.nonempty) : + ∃ C₀ C₁ : set α, (perfect C₀ ∧ C₀.nonempty ∧ C₀ ⊆ C) ∧ + (perfect C₁ ∧ C₁.nonempty ∧ C₁ ⊆ C) ∧ disjoint C₀ C₁ := +begin + cases hnonempty with y yC, + obtain ⟨x, xC, hxy⟩ : ∃ x ∈ C, x ≠ y, + { have := hC.acc _ yC, + rw acc_pt_iff_nhds at this, + rcases this univ (univ_mem) with ⟨x, xC, hxy⟩, + exact ⟨x, xC.2, hxy⟩, }, + obtain ⟨U, xU, Uop, V, yV, Vop, hUV⟩ := exists_open_nhds_disjoint_closure hxy, + use [closure (U ∩ C), closure (V ∩ C)], + split; rw ← and_assoc, + { refine ⟨hC.closure_nhds_inter x xC xU Uop, _⟩, + rw hC.closed.closure_subset_iff, + exact inter_subset_right _ _, }, + split, + { refine ⟨hC.closure_nhds_inter y yC yV Vop, _⟩, + rw hC.closed.closure_subset_iff, + exact inter_subset_right _ _, }, + apply disjoint.mono _ _ hUV; apply closure_mono; exact inter_subset_left _ _, +end + +section kernel + +/-- The **Cantor-Bendixson Theorem**: Any closed subset of a second countable space +can be written as the union of a countable set and a perfect set.-/ +theorem exists_countable_union_perfect_of_is_closed [second_countable_topology α] + (hclosed : is_closed C) : + ∃ V D : set α, (V.countable) ∧ (perfect D) ∧ (C = V ∪ D) := +begin + obtain ⟨b, bct, bnontrivial, bbasis⟩ := topological_space.exists_countable_basis α, + let v := {U ∈ b | (U ∩ C).countable}, + let V := ⋃ U ∈ v, U, + let D := C \ V, + have Vct : (V ∩ C).countable, + { simp only [Union_inter, mem_sep_iff], + apply countable.bUnion, + { exact countable.mono (inter_subset_left _ _) bct, }, + { exact inter_subset_right _ _, }, }, + refine ⟨V ∩ C, D, Vct, ⟨_, _⟩, _⟩, + { refine hclosed.sdiff (is_open_bUnion (λ U, _)), + exact λ ⟨Ub, _⟩, is_topological_basis.is_open bbasis Ub, }, + { rw preperfect_iff_nhds, + intros x xD E xE, + have : ¬ (E ∩ D).countable, + { intro h, + obtain ⟨U, hUb, xU, hU⟩ : ∃ U ∈ b, x ∈ U ∧ U ⊆ E, + { exact (is_topological_basis.mem_nhds_iff bbasis).mp xE, }, + have hU_cnt : (U ∩ C).countable, + { apply @countable.mono _ _ ((E ∩ D) ∪ (V ∩ C)), + { rintros y ⟨yU, yC⟩, + by_cases y ∈ V, + { exact mem_union_right _ (mem_inter h yC), }, + { exact mem_union_left _ (mem_inter (hU yU) ⟨yC, h⟩), }, }, + exact countable.union h Vct, }, + have : U ∈ v := ⟨hUb, hU_cnt⟩, + apply xD.2, + exact mem_bUnion this xU, }, + by_contradiction h, + push_neg at h, + exact absurd (countable.mono h (set.countable_singleton _)) this, }, + { rw [inter_comm, inter_union_diff], }, +end + +/-- Any uncountable closed set in a second countable space contains a nonempty perfect subset.-/ +theorem exists_perfect_nonempty_of_is_closed_of_not_countable [second_countable_topology α] + (hclosed : is_closed C) (hunc : ¬ C.countable) : + ∃ D : set α, perfect D ∧ D.nonempty ∧ D ⊆ C := +begin + rcases exists_countable_union_perfect_of_is_closed hclosed with ⟨V, D, Vct, Dperf, VD⟩, + refine ⟨D, ⟨Dperf, _⟩⟩, + split, + { rw nonempty_iff_ne_empty, + by_contradiction, + rw [h, union_empty] at VD, + rw VD at hunc, + contradiction, }, + rw VD, + exact subset_union_right _ _, +end + +end kernel From d4f69d96f3532729da8ebb763f4bc26fcf640f06 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sun, 18 Dec 2022 18:49:52 +0000 Subject: [PATCH 0074/1291] feat(algebra/big_operators/basic): Sum of `ite` (#16825) A sum of if then else that don't happen simultaneously can be written as a single if then else. --- src/algebra/big_operators/basic.lean | 15 +++++++++++++++ src/data/set/pairwise.lean | 27 +++++++++++++++++++++++++++ 2 files changed, 42 insertions(+) diff --git a/src/algebra/big_operators/basic.lean b/src/algebra/big_operators/basic.lean index 98797855b4e12..b8641c0685d44 100644 --- a/src/algebra/big_operators/basic.lean +++ b/src/algebra/big_operators/basic.lean @@ -1328,6 +1328,21 @@ begin rwa eq_of_mem_of_not_mem_erase hx hnx end +/-- See also `finset.prod_boole`. -/ +@[to_additive "See also `finset.sum_boole`."] +lemma prod_ite_one {f : α → Prop} [decidable_pred f] (hf : (s : set α).pairwise_disjoint f) + (a : β) : + ∏ i in s, ite (f i) a 1 = ite (∃ i ∈ s, f i) a 1 := +begin + split_ifs, + { obtain ⟨i, hi, hfi⟩ := h, + rw [prod_eq_single_of_mem _ hi, if_pos hfi], + exact λ j hj h, if_neg (λ hfj, (hf hj hi h).le_bot ⟨hfj, hfi⟩) }, + { push_neg at h, + rw prod_eq_one, + exact λ i hi, if_neg (h i hi) } +end + lemma sum_erase_lt_of_pos {γ : Type*} [decidable_eq α] [ordered_add_comm_monoid γ] [covariant_class γ γ (+) (<)] {s : finset α} {d : α} (hd : d ∈ s) {f : α → γ} (hdf : 0 < f d) : ∑ (m : α) in s.erase d, f m < ∑ (m : α) in s, f m := diff --git a/src/data/set/pairwise.lean b/src/data/set/pairwise.lean index 251ba470614af..ba87b6e47e941 100644 --- a/src/data/set/pairwise.lean +++ b/src/data/set/pairwise.lean @@ -136,19 +136,37 @@ lemma pairwise_insert : by simp only [insert_eq, pairwise_union, pairwise_singleton, true_and, mem_singleton_iff, forall_eq] +lemma pairwise_insert_of_not_mem (ha : a ∉ s) : + (insert a s).pairwise r ↔ s.pairwise r ∧ ∀ b ∈ s, r a b ∧ r b a := +pairwise_insert.trans $ and_congr_right' $ forall₂_congr $ λ b hb, + by simp [(ne_of_mem_of_not_mem hb ha).symm] + lemma pairwise.insert (hs : s.pairwise r) (h : ∀ b ∈ s, a ≠ b → r a b ∧ r b a) : (insert a s).pairwise r := pairwise_insert.2 ⟨hs, h⟩ +lemma pairwise.insert_of_not_mem (ha : a ∉ s) (hs : s.pairwise r) (h : ∀ b ∈ s, r a b ∧ r b a) : + (insert a s).pairwise r := +(pairwise_insert_of_not_mem ha).2 ⟨hs, h⟩ + lemma pairwise_insert_of_symmetric (hr : symmetric r) : (insert a s).pairwise r ↔ s.pairwise r ∧ ∀ b ∈ s, a ≠ b → r a b := by simp only [pairwise_insert, hr.iff a, and_self] +lemma pairwise_insert_of_symmetric_of_not_mem (hr : symmetric r) (ha : a ∉ s) : + (insert a s).pairwise r ↔ s.pairwise r ∧ ∀ b ∈ s, r a b := +by simp only [pairwise_insert_of_not_mem ha, hr.iff a, and_self] + lemma pairwise.insert_of_symmetric (hs : s.pairwise r) (hr : symmetric r) (h : ∀ b ∈ s, a ≠ b → r a b) : (insert a s).pairwise r := (pairwise_insert_of_symmetric hr).2 ⟨hs, h⟩ +lemma pairwise.insert_of_symmetric_of_not_mem (hs : s.pairwise r) (hr : symmetric r) (ha : a ∉ s) + (h : ∀ b ∈ s, r a b) : + (insert a s).pairwise r := +(pairwise_insert_of_symmetric_of_not_mem hr ha).2 ⟨hs, h⟩ + lemma pairwise_pair : set.pairwise {a, b} r ↔ (a ≠ b → r a b ∧ r b a) := by simp [pairwise_insert] @@ -226,11 +244,20 @@ lemma pairwise_disjoint_insert {i : ι} : ↔ s.pairwise_disjoint f ∧ ∀ j ∈ s, i ≠ j → disjoint (f i) (f j) := set.pairwise_insert_of_symmetric $ symmetric_disjoint.comap f +lemma pairwise_disjoint_insert_of_not_mem {i : ι} (hi : i ∉ s) : + (insert i s).pairwise_disjoint f ↔ s.pairwise_disjoint f ∧ ∀ j ∈ s, disjoint (f i) (f j) := +pairwise_insert_of_symmetric_of_not_mem (symmetric_disjoint.comap f) hi + lemma pairwise_disjoint.insert (hs : s.pairwise_disjoint f) {i : ι} (h : ∀ j ∈ s, i ≠ j → disjoint (f i) (f j)) : (insert i s).pairwise_disjoint f := set.pairwise_disjoint_insert.2 ⟨hs, h⟩ +lemma pairwise_disjoint.insert_of_not_mem (hs : s.pairwise_disjoint f) {i : ι} (hi : i ∉ s) + (h : ∀ j ∈ s, disjoint (f i) (f j)) : + (insert i s).pairwise_disjoint f := +(set.pairwise_disjoint_insert_of_not_mem hi).2 ⟨hs, h⟩ + lemma pairwise_disjoint.image_of_le (hs : s.pairwise_disjoint f) {g : ι → ι} (hg : f ∘ g ≤ f) : (g '' s).pairwise_disjoint f := begin From d2617e08382c6f606d78efff82ce2e21b408fd6c Mon Sep 17 00:00:00 2001 From: Lukas Miaskiwskyi Date: Mon, 19 Dec 2022 01:08:49 +0000 Subject: [PATCH 0075/1291] feat(analysis/locally_convex): locally convex hausdorff spaces (#17937) Adds that locally convex spaces whose topology is induced by a family of seminorms are Hausdorff iff the family of seminorms is separating, together with some helper lemmas. Closes issue #15565. Co-authored-by: Lukas Miaskiwskyi --- .../locally_convex/with_seminorms.lean | 80 +++++++++++++++++++ src/analysis/seminorm.lean | 21 +++++ 2 files changed, 101 insertions(+) diff --git a/src/analysis/locally_convex/with_seminorms.lean b/src/analysis/locally_convex/with_seminorms.lean index 6fb325e9d33b3..0294d199c1771 100644 --- a/src/analysis/locally_convex/with_seminorms.lean +++ b/src/analysis/locally_convex/with_seminorms.lean @@ -283,12 +283,92 @@ lemma with_seminorms.with_seminorms_eq {p : seminorm_family 𝕜 E ι} [t : topo variables [topological_space E] variables {p : seminorm_family 𝕜 E ι} +lemma with_seminorms.topological_add_group (hp : with_seminorms p) : topological_add_group E := +begin + rw hp.with_seminorms_eq, + exact add_group_filter_basis.is_topological_add_group _ +end + lemma with_seminorms.has_basis (hp : with_seminorms p) : (𝓝 (0 : E)).has_basis (λ (s : set E), s ∈ p.basis_sets) id := begin rw (congr_fun (congr_arg (@nhds E) hp.1) 0), exact add_group_filter_basis.nhds_zero_has_basis _, end + +lemma with_seminorms.has_basis_zero_ball (hp : with_seminorms p) : (𝓝 (0 : E)).has_basis + (λ sr : finset ι × ℝ, 0 < sr.2) (λ sr, (sr.1.sup p).ball 0 sr.2) := +begin + refine ⟨λ V, _⟩, + simp only [hp.has_basis.mem_iff, seminorm_family.basis_sets_iff, prod.exists], + split, + { rintros ⟨-, ⟨s, r, hr, rfl⟩, hV⟩, + exact ⟨s, r, hr, hV⟩ }, + { rintros ⟨s, r, hr, hV⟩, + exact ⟨_, ⟨s, r, hr, rfl⟩, hV⟩ } +end + +lemma with_seminorms.has_basis_ball (hp : with_seminorms p) {x : E} : (𝓝 (x : E)).has_basis + (λ sr : finset ι × ℝ, 0 < sr.2) (λ sr, (sr.1.sup p).ball x sr.2) := +begin + haveI : topological_add_group E := hp.topological_add_group, + rw [← map_add_left_nhds_zero], + convert (hp.has_basis_zero_ball.map ((+) x)), + ext sr : 1, + have : (sr.fst.sup p).ball (x +ᵥ 0) sr.snd = x +ᵥ (sr.fst.sup p).ball 0 sr.snd + := eq.symm (seminorm.vadd_ball (sr.fst.sup p)), + rwa [vadd_eq_add, add_zero] at this, +end + +/-- The `x`-neighbourhoods of a space whose topology is induced by a family of seminorms +are exactly the sets which contain seminorm balls around `x`.-/ +lemma with_seminorms.mem_nhds_iff (hp : with_seminorms p) (x : E) (U : set E) : + U ∈ nhds x ↔ ∃ (s : finset ι) (r > 0), (s.sup p).ball x r ⊆ U := +by rw [hp.has_basis_ball.mem_iff, prod.exists] + +/-- The open sets of a space whose topology is induced by a family of seminorms +are exactly the sets which contain seminorm balls around all of their points.-/ +lemma with_seminorms.is_open_iff_mem_balls (hp : with_seminorms p) (U : set E) : + is_open U ↔ ∀ x ∈ U, ∃ (s : finset ι) (r > 0), (s.sup p).ball x r ⊆ U := +by simp_rw [←with_seminorms.mem_nhds_iff hp _ U, is_open_iff_mem_nhds] + +/- Note that through the following lemmas, one also immediately has that separating families +of seminorms induce T₂ and T₃ topologies by `topological_add_group.t2_space` +and `topological_add_group.t3_space` -/ +/-- A separating family of seminorms induces a T₁ topology. -/ +lemma with_seminorms.t1_of_separating (hp : with_seminorms p) (h : ∀ x ≠ 0, ∃ i, p i x ≠ 0) : + t1_space E := +begin + haveI := hp.topological_add_group, + refine topological_add_group.t1_space _ _, + rw [← is_open_compl_iff, hp.is_open_iff_mem_balls], + rintros x (hx : x ≠ 0), + cases h x hx with i pi_nonzero, + refine ⟨{i}, p i x, by positivity, subset_compl_singleton_iff.mpr _⟩, + rw [finset.sup_singleton, mem_ball, zero_sub, map_neg_eq_map, not_lt] +end + +/-- A family of seminorms inducing a T₁ topology is separating. -/ +lemma with_seminorms.separating_of_t1 [t1_space E] (hp : with_seminorms p) (x : E) (hx : x ≠ 0) : + ∃ i, p i x ≠ 0 := +begin + have := ((t1_space_tfae E).out 0 9).mp infer_instance, + by_contra' h, + refine hx (this _), + rw hp.has_basis_zero_ball.specializes_iff, + rintros ⟨s, r⟩ (hr : 0 < r), + simp only [ball_finset_sup_eq_Inter _ _ _ hr, mem_Inter₂, mem_ball_zero, h, hr, forall_true_iff], +end + +/-- A family of seminorms is separating iff it induces a T₁ topology. -/ +lemma with_seminorms.separating_iff_t1 (hp : with_seminorms p) : + (∀ x ≠ 0, ∃ i, p i x ≠ 0) ↔ t1_space E := +begin + refine ⟨with_seminorms.t1_of_separating hp, _⟩, + introI, + exact with_seminorms.separating_of_t1 hp, +end + end topology section topological_add_group diff --git a/src/analysis/seminorm.lean b/src/analysis/seminorm.lean index a06e236385dc4..1fce620c351c3 100644 --- a/src/analysis/seminorm.lean +++ b/src/analysis/seminorm.lean @@ -6,6 +6,7 @@ Authors: Jean Lo, Yaël Dillies, Moritz Doll import data.real.pointwise import analysis.convex.function import analysis.locally_convex.basic +import analysis.normed.group.add_torsor /-! # Seminorms @@ -609,6 +610,26 @@ begin exact (map_add_le_add p _ _).trans (add_le_add hy₁ hy₂) end +lemma sub_mem_ball (p : seminorm 𝕜 E) (x₁ x₂ y : E) (r : ℝ) : + x₁ - x₂ ∈ p.ball y r ↔ x₁ ∈ p.ball (x₂ + y) r := +by simp_rw [mem_ball, sub_sub] + +/-- The image of a ball under addition with a singleton is another ball. -/ +lemma vadd_ball (p : seminorm 𝕜 E) : + x +ᵥ p.ball y r = p.ball (x +ᵥ y) r := +begin + letI := add_group_seminorm.to_seminormed_add_comm_group p.to_add_group_seminorm, + exact vadd_ball x y r, +end + +/-- The image of a closed ball under addition with a singleton is another closed ball. -/ +lemma vadd_closed_ball (p : seminorm 𝕜 E) : + x +ᵥ p.closed_ball y r = p.closed_ball (x +ᵥ y) r := +begin + letI := add_group_seminorm.to_seminormed_add_comm_group p.to_add_group_seminorm, + exact vadd_closed_ball x y r, +end + end has_smul section module From bbeb185db4ccee8ed07dc48449414ebfa39cb821 Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Mon, 19 Dec 2022 06:58:15 +0000 Subject: [PATCH 0076/1291] chore(data/polynomial/derivative): replace n by C n (#17911) Rename `derivative_pow` lemmas, refactor `derivative_X_add_pow` slightly, and replace `n` by `C n` analogously to #17795 --- src/data/polynomial/derivative.lean | 61 +++++++++++-------- src/data/polynomial/expand.lean | 2 +- src/data/polynomial/field_division.lean | 4 +- src/field_theory/finite/galois_field.lean | 2 +- .../cyclotomic/discriminant.lean | 6 +- src/ring_theory/polynomial/bernstein.lean | 7 ++- 6 files changed, 48 insertions(+), 34 deletions(-) diff --git a/src/data/polynomial/derivative.lean b/src/data/polynomial/derivative.lean index cc33b32a29fed..170c4700a4fba 100644 --- a/src/data/polynomial/derivative.lean +++ b/src/data/polynomial/derivative.lean @@ -108,6 +108,9 @@ by simp [bit1] @[simp] lemma derivative_add {f g : R[X]} : derivative (f + g) = derivative f + derivative g := derivative.map_add f g +@[simp] lemma derivative_X_add_C (c : R) : (X + C c).derivative = 1 := +by rw [derivative_add, derivative_X, derivative_C, add_zero] + @[simp] lemma iterate_derivative_add {f g : R[X]} {k : ℕ} : derivative^[k] (f + g) = (derivative^[k] f) + (derivative^[k] g) := derivative.to_add_monoid_hom.iterate_map_add _ _ _ @@ -391,22 +394,25 @@ section comm_semiring variables [comm_semiring R] theorem derivative_pow_succ (p : R[X]) (n : ℕ) : - (p ^ (n + 1)).derivative = (n + 1) * (p ^ n) * p.derivative := -nat.rec_on n (by rw [pow_one, nat.cast_zero, zero_add, one_mul, pow_zero, one_mul]) $ λ n ih, -by rw [pow_succ', derivative_mul, ih, mul_right_comm, ← add_mul, - add_mul (n.succ : R[X]), one_mul, pow_succ', mul_assoc, n.cast_succ] + (p ^ (n + 1)).derivative = C ↑(n + 1) * (p ^ n) * p.derivative := +nat.rec_on n (by rw [pow_one, nat.cast_one, C_1, one_mul, pow_zero, one_mul]) $ λ n ih, +by rw [pow_succ', derivative_mul, ih, nat.add_one, mul_right_comm, nat.cast_add n.succ, C_add, + add_mul, add_mul, pow_succ', ← mul_assoc, nat.cast_one, C_1, one_mul] theorem derivative_pow (p : R[X]) (n : ℕ) : - (p ^ n).derivative = n * (p ^ (n - 1)) * p.derivative := -nat.cases_on n (by rw [pow_zero, derivative_one, nat.cast_zero, zero_mul, zero_mul]) $ λ n, + (p ^ n).derivative = C ↑n * (p ^ (n - 1)) * p.derivative := +nat.cases_on n (by rw [pow_zero, derivative_one, nat.cast_zero, C_0, zero_mul, zero_mul]) $ λ n, by rw [p.derivative_pow_succ n, n.succ_sub_one, n.cast_succ] +theorem derivative_sq (p : R[X]) : (p ^ 2).derivative = C 2 * p * p.derivative := +by rw [derivative_pow_succ, nat.cast_two, pow_one] + theorem dvd_iterate_derivative_pow (f : R[X]) (n : ℕ) {m : ℕ} (c : R) (hm : m ≠ 0) : (n : R) ∣ eval c (derivative^[m] (f ^ n)) := begin obtain ⟨m, rfl⟩ := nat.exists_eq_succ_of_ne_zero hm, - rw [function.iterate_succ_apply, derivative_pow, mul_assoc, iterate_derivative_nat_cast_mul, - eval_mul, eval_nat_cast], + rw [function.iterate_succ_apply, derivative_pow, mul_assoc, C_eq_nat_cast, + iterate_derivative_nat_cast_mul, eval_mul, eval_nat_cast], exact dvd_mul_right _ _, end @@ -428,11 +434,14 @@ lemma iterate_derivative_X_pow_eq_smul (n : ℕ) (k : ℕ) : (derivative^[k] (X ^ n : R[X])) = (nat.desc_factorial n k : R) • X ^ (n - k) := by rw [iterate_derivative_X_pow_eq_C_mul n k, smul_eq_C_mul] -lemma derivative_X_add_pow (c : R) (m : ℕ) : ((X + C c) ^ m).derivative = m * (X + C c) ^ (m - 1) := -by rw [derivative_pow, derivative_add, derivative_X, derivative_C, add_zero, mul_one] +lemma derivative_X_add_C_pow (c : R) (m : ℕ) : + ((X + C c) ^ m).derivative = C ↑m * (X + C c) ^ (m - 1) := +by rw [derivative_pow, derivative_X_add_C, mul_one] + +lemma derivative_X_add_C_sq (c : R) : ((X + C c) ^ 2).derivative = C 2 * (X + C c) := +by rw [derivative_sq, derivative_X_add_C, mul_one] -lemma iterate_derivative_X_add_pow (n k : ℕ) (c : R) : - (derivative^[k] ((X + C c) ^ n)) = +lemma iterate_derivative_X_add_pow (n k : ℕ) (c : R) : derivative^[k] ((X + C c) ^ n) = ↑(∏ i in finset.range k, (n - i)) * (X + C c) ^ (n - k) := begin induction k with k IH, @@ -440,11 +449,10 @@ begin tsub_zero] }, { simp only [function.iterate_succ_apply', IH, derivative_mul, zero_mul, derivative_nat_cast, zero_add, finset.prod_range_succ, C_eq_nat_cast, nat.sub_sub, ←mul_assoc, - derivative_X_add_pow, nat.succ_eq_add_one, nat.cast_mul] }, + derivative_X_add_C_pow, nat.succ_eq_add_one, nat.cast_mul] }, end -lemma derivative_comp (p q : R[X]) : - (p.comp q).derivative = q.derivative * p.derivative.comp q := +lemma derivative_comp (p q : R[X]) : (p.comp q).derivative = q.derivative * p.derivative.comp q := begin apply polynomial.induction_on' p, { intros p₁ p₂ h₁ h₂, simp [h₁, h₂, mul_add], }, @@ -497,10 +505,12 @@ linear_map.map_neg derivative f derivative^[k] (-f) = - (derivative^[k] f) := (@derivative R _).to_add_monoid_hom.iterate_map_neg _ _ -@[simp] lemma derivative_sub {f g : R[X]} : - derivative (f - g) = derivative f - derivative g := +@[simp] lemma derivative_sub {f g : R[X]} : derivative (f - g) = derivative f - derivative g := linear_map.map_sub derivative f g +@[simp] lemma derivative_X_sub_C (c : R) : (X - C c).derivative = 1 := +by rw [derivative_sub, derivative_X, derivative_C, sub_zero] + @[simp] lemma iterate_derivative_sub {k : ℕ} {f g : R[X]} : derivative^[k] (f - g) = (derivative^[k] f) - (derivative^[k] g) := by induction k with k ih generalizing f g; simp* @@ -525,12 +535,12 @@ section comm_ring variables [comm_ring R] lemma derivative_comp_one_sub_X (p : R[X]) : - (p.comp (1-X)).derivative = -p.derivative.comp (1-X) := + (p.comp (1 - X)).derivative = -p.derivative.comp (1 - X) := by simp [derivative_comp] @[simp] lemma iterate_derivative_comp_one_sub_X (p : R[X]) (k : ℕ) : - derivative^[k] (p.comp (1-X)) = (-1)^k * (derivative^[k] p).comp (1-X) := + derivative^[k] (p.comp (1 - X)) = (-1) ^ k * (derivative^[k] p).comp (1 - X) := begin induction k with k ih generalizing p, { simp, }, @@ -545,16 +555,19 @@ begin simpa using (eval_ring_hom r).map_multiset_prod (multiset.map (λ a, X - C a) (S.erase r)), end -lemma derivative_X_sub_pow (c : R) (m : ℕ) : - ((X - C c) ^ m).derivative = m * (X - C c) ^ (m - 1) := -by rw [derivative_pow, derivative_sub, derivative_X, derivative_C, sub_zero, mul_one] +lemma derivative_X_sub_C_pow (c : R) (m : ℕ) : + ((X - C c) ^ m).derivative = C ↑m * (X - C c) ^ (m - 1) := +by rw [derivative_pow, derivative_X_sub_C, mul_one] + +lemma derivative_X_sub_C_sq (c : R) : ((X - C c) ^ 2).derivative = C 2 * (X - C c) := +by rw [derivative_sq, derivative_X_sub_C, mul_one] lemma iterate_derivative_X_sub_pow (n k : ℕ) (c : R) : - (derivative^[k] ((X - C c) ^ n)) = - (↑(∏ i in finset.range k, (n - i))) * (X - C c) ^ (n - k) := + (derivative^[k] ((X - C c) ^ n)) = (↑(∏ i in finset.range k, (n - i))) * (X - C c) ^ (n - k) := by simp_rw [sub_eq_add_neg, ←C_neg, iterate_derivative_X_add_pow] end comm_ring end derivative + end polynomial diff --git a/src/data/polynomial/expand.lean b/src/data/polynomial/expand.lean index b10b2f384b09d..0ea5919000489 100644 --- a/src/data/polynomial/expand.lean +++ b/src/data/polynomial/expand.lean @@ -70,7 +70,7 @@ by rw [function.iterate_succ_apply', pow_succ, expand_mul, ih] theorem derivative_expand (f : R[X]) : (expand R p f).derivative = expand R p f.derivative * (p * X ^ (p - 1)) := -by rw [coe_expand, derivative_eval₂_C, derivative_pow, derivative_X, mul_one] +by rw [coe_expand, derivative_eval₂_C, derivative_pow, C_eq_nat_cast, derivative_X, mul_one] theorem coeff_expand {p : ℕ} (hp : 0 < p) (f : R[X]) (n : ℕ) : (expand R p f).coeff n = if p ∣ n then f.coeff (n / p) else 0 := diff --git a/src/data/polynomial/field_division.lean b/src/data/polynomial/field_division.lean index 9fc76e6814d0b..b1a6db3f07905 100644 --- a/src/data/polynomial/field_division.lean +++ b/src/data/polynomial/field_division.lean @@ -36,12 +36,12 @@ begin have hn : n + 1 = _ := tsub_add_cancel_of_le ((root_multiplicity_pos hp).mpr hpt), rw ←hn, set q := p /ₘ (X - C t) ^ (n + 1) with hq, - convert_to root_multiplicity t ((X - C t) ^ n * (derivative q * (X - C t) + q * ↑(n + 1))) = n, + convert_to root_multiplicity t ((X - C t) ^ n * (derivative q * (X - C t) + q * C ↑(n + 1))) = n, { congr, rw [mul_add, mul_left_comm $ (X - C t) ^ n, ←pow_succ'], congr' 1, rw [mul_left_comm $ (X - C t) ^ n, mul_comm $ (X - C t) ^ n] }, - have h : (derivative q * (X - C t) + q * ↑(n + 1)).eval t ≠ 0, + have h : (derivative q * (X - C t) + q * C ↑(n + 1)).eval t ≠ 0, { suffices : eval t q * ↑(n + 1) ≠ 0, { simpa }, refine mul_ne_zero _ (nat.cast_ne_zero.mpr n.succ_ne_zero), diff --git a/src/field_theory/finite/galois_field.lean b/src/field_theory/finite/galois_field.lean index b82c540a7203d..ac9250227e85c 100644 --- a/src/field_theory/finite/galois_field.lean +++ b/src/field_theory/finite/galois_field.lean @@ -38,7 +38,7 @@ lemma galois_poly_separable {K : Type*} [field K] (p q : ℕ) [char_p K p] (h : begin use [1, (X ^ q - X - 1)], rw [← char_p.cast_eq_zero_iff K[X] p] at h, - rw [derivative_sub, derivative_pow, derivative_X, h], + rw [derivative_sub, derivative_X_pow, derivative_X, C_eq_nat_cast, h], ring, end diff --git a/src/number_theory/cyclotomic/discriminant.lean b/src/number_theory/cyclotomic/discriminant.lean index 61de05a851897..f268b04516fcd 100644 --- a/src/number_theory/cyclotomic/discriminant.lean +++ b/src/number_theory/cyclotomic/discriminant.lean @@ -103,9 +103,9 @@ begin hp.out.one_lt) (pow_pos hp.out.pos _))) (even.mul_right (nat.even_sub_one_of_prime_ne_two hp.out hptwo) _) odd_one) } }, { have H := congr_arg derivative (cyclotomic_prime_pow_mul_X_pow_sub_one K p k), - rw [derivative_mul, derivative_sub, derivative_one, sub_zero, derivative_pow, - derivative_X, mul_one, derivative_sub, derivative_one, sub_zero, derivative_pow, - derivative_X, mul_one, ← pnat.pow_coe, hζ.minpoly_eq_cyclotomic_of_irreducible hirr] at H, + rw [derivative_mul, derivative_sub, derivative_one, sub_zero, derivative_X_pow, C_eq_nat_cast, + derivative_sub, derivative_one, sub_zero, derivative_X_pow, C_eq_nat_cast, ← pnat.pow_coe, + hζ.minpoly_eq_cyclotomic_of_irreducible hirr] at H, replace H := congr_arg (λ P, aeval ζ P) H, simp only [aeval_add, aeval_mul, minpoly.aeval, zero_mul, add_zero, aeval_nat_cast, _root_.map_sub, aeval_one, aeval_X_pow] at H, diff --git a/src/ring_theory/polynomial/bernstein.lean b/src/ring_theory/polynomial/bernstein.lean index 6a6aecf3fc47a..cd545d889dc7f 100644 --- a/src/ring_theory/polynomial/bernstein.lean +++ b/src/ring_theory/polynomial/bernstein.lean @@ -98,14 +98,15 @@ lemma derivative_succ_aux (n ν : ℕ) : begin rw [bernstein_polynomial], suffices : - ↑((n + 1).choose (ν + 1)) * ((↑ν + 1) * X ^ ν) * (1 - X) ^ (n - ν) + ↑((n + 1).choose (ν + 1)) * (↑(ν + 1) * X ^ ν) * (1 - X) ^ (n - ν) -(↑((n + 1).choose (ν + 1)) * X ^ (ν + 1) * (↑(n - ν) * (1 - X) ^ (n - ν - 1))) = - (↑n + 1) * (↑(n.choose ν) * X ^ ν * (1 - X) ^ (n - ν) - + ↑(n + 1) * (↑(n.choose ν) * X ^ ν * (1 - X) ^ (n - ν) - ↑(n.choose (ν + 1)) * X ^ (ν + 1) * (1 - X) ^ (n - (ν + 1))), { simpa only [polynomial.derivative_pow, ←sub_eq_add_neg, nat.succ_sub_succ_eq_sub, polynomial.derivative_mul, polynomial.derivative_nat_cast, zero_mul, nat.cast_add, algebra_map.coe_one, polynomial.derivative_X, mul_one, zero_add, - polynomial.derivative_sub, polynomial.derivative_one, zero_sub, mul_neg] }, + polynomial.derivative_sub, polynomial.derivative_one, zero_sub, mul_neg, + nat.sub_zero, ← nat.cast_succ, polynomial.C_eq_nat_cast], }, conv_rhs { rw mul_sub, }, -- We'll prove the two terms match up separately. refine congr (congr_arg has_sub.sub _) _, From b7487a0ea553aa8d9578aa3af263aaf521686c54 Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Mon, 19 Dec 2022 09:05:51 +0000 Subject: [PATCH 0077/1291] feat(algebraic_geometry/EllipticCurve/weierstrass): define Weierstrass polynomials and prove basic facts (#17631) Define the Weierstrass polynomial and its partial derivatives, as well as properties of their evaluations (the Weierstrass equation and nonsingularity). Prove that a Weierstrass curve is nonsingular at every point if its discriminant is non-zero, and its coordinate ring is an integral domain (because the associated polynomial is irreducible). Fix minor issues (rename the variable `C` with the variable `W` to avoid a clash with `polynomial.C`, and generalise `two_torsion_polynomial_disc_ne_zero`). Co-authored-by: Junyan Xu --- docs/references.bib | 8 + .../elliptic_curve/weierstrass.lean | 264 ++++++++++++++---- 2 files changed, 224 insertions(+), 48 deletions(-) diff --git a/docs/references.bib b/docs/references.bib index 84b6f8ab0b2a8..8620c56acc625 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -1875,6 +1875,14 @@ @Book{ serre1965 url = {https://doi.org/10.1007/978-1-4757-3910-7} } +@Book{ silverman2009, + author = {Silverman, Joseph}, + publisher = {Springer New York, NY}, + series = {Graduate Texts in Mathematics}, + title = {The Arithmetic of Elliptic Curves}, + year = {2009} +} + @Book{ simon2011, author = {Simon, Barry}, title = {Convexity: An Analytic Viewpoint}, diff --git a/src/algebraic_geometry/elliptic_curve/weierstrass.lean b/src/algebraic_geometry/elliptic_curve/weierstrass.lean index 750dd897746c6..78a0cec827478 100644 --- a/src/algebraic_geometry/elliptic_curve/weierstrass.lean +++ b/src/algebraic_geometry/elliptic_curve/weierstrass.lean @@ -5,12 +5,13 @@ Authors: Kevin Buzzard, David Kurniadi Angdinata -/ import algebra.cubic_discriminant +import ring_theory.adjoin_root import tactic.linear_combination /-! # Weierstrass equations of elliptic curves -We give a working definition of an elliptic curve as a non-singular Weierstrass curve given by a +We give a working definition of an elliptic curve as a nonsingular Weierstrass curve given by a Weierstrass equation, which is mathematically accurate in many cases but also good for computation. ## Mathematical background @@ -21,17 +22,23 @@ is smooth and proper and the fibres are geometrically-connected one-dimensional the special case where `S` is the spectrum of some commutative ring `R` whose Picard group is zero (this includes all fields, all PIDs, and many other commutative rings) it can be shown (using a lot of algebro-geometric machinery) that every elliptic curve `E` is a projective plane cubic isomorphic -to a Weierstrass curve given by the equation $y^2 + a_1xy + a_3y = x^3 + a_2x^2 + a_4x + a_6$ for +to a Weierstrass curve given by the equation $Y^2 + a_1XY + a_3Y = X^3 + a_2X^2 + a_4X + a_6$ for some $a_i$ in `R`, and such that a certain quantity called the discriminant of `E` is a unit in `R`. If `R` is a field, this quantity divides the discriminant of a cubic polynomial whose roots over a -splitting field of `R` are precisely the x-coordinates of the non-zero 2-torsion points of `E`. +splitting field of `R` are precisely the $X$-coordinates of the non-zero 2-torsion points of `E`. ## Main definitions * `weierstrass_curve`: a Weierstrass curve over a commutative ring. + * `weierstrass_curve.Δ`: the discriminant of a Weierstrass curve. * `weierstrass_curve.variable_change`: the Weierstrass curve induced by a change of variables. * `weierstrass_curve.base_change`: the Weierstrass curve base changed over an algebra. * `weierstrass_curve.two_torsion_polynomial`: the 2-torsion polynomial of a Weierstrass curve. + * `weierstrass_curve.polynomial`: the polynomial associated to a Weierstrass curve. + * `weierstrass_curve.equation`: the Weirstrass equation of a Weierstrass curve. + * `weierstrass_curve.nonsingular`: the nonsingular condition at a point on a Weierstrass curve. + * `weierstrass_curve.coordinate_ring`: the coordinate ring of a Weierstrass curve. + * `weierstrass_curve.function_field`: the function field of a Weierstrass curve. * `elliptic_curve`: an elliptic curve over a commutative ring. * `elliptic_curve.j`: the j-invariant of an elliptic curve. @@ -39,6 +46,11 @@ splitting field of `R` are precisely the x-coordinates of the non-zero 2-torsion * `weierstrass_curve.two_torsion_polynomial_disc`: the discriminant of a Weierstrass curve is a constant factor of the cubic discriminant of its 2-torsion polynomial. + * `weierstrass_curve.nonsingular_of_Δ_ne_zero`: a Weierstrass curve is nonsingular at every point + if its discriminant is non-zero. + * `weierstrass_curve.coordinate_ring.is_domain`: the coordinate ring of a Weierstrass curve is + an integral domain. + * `elliptic_curve.nonsingular`: an elliptic curve is nonsingular at every point. * `elliptic_curve.variable_change_j`: the j-invariant of an elliptic curve is invariant under an admissible linear change of variables. @@ -65,13 +77,16 @@ elliptic curve, weierstrass equation, j invariant private meta def map_simp : tactic unit := `[simp only [map_one, map_bit0, map_bit1, map_neg, map_add, map_sub, map_mul, map_pow]] +private meta def eval_simp : tactic unit := +`[simp only [eval_C, eval_X, eval_add, eval_sub, eval_mul, eval_pow]] + universes u v variable {R : Type u} /-! ## Weierstrass curves -/ -/-- A Weierstrass curve $y^2 + a_1xy + a_3y = x^3 + a_2x^2 + a_4x + a_6$ with parameters $a_i$. -/ +/-- A Weierstrass curve $Y^2 + a_1XY + a_3Y = X^3 + a_2X^2 + a_4X + a_6$ with parameters $a_i$. -/ @[ext] structure weierstrass_curve (R : Type u) := (a₁ a₂ a₃ a₄ a₆ : R) instance [inhabited R] : inhabited $ weierstrass_curve R := @@ -79,40 +94,40 @@ instance [inhabited R] : inhabited $ weierstrass_curve R := namespace weierstrass_curve -variables [comm_ring R] (C : weierstrass_curve R) +variables [comm_ring R] (W : weierstrass_curve R) section quantity /-! ### Standard quantities -/ /-- The `b₂` coefficient of a Weierstrass curve. -/ -@[simp] def b₂ : R := C.a₁ ^ 2 + 4 * C.a₂ +@[simp] def b₂ : R := W.a₁ ^ 2 + 4 * W.a₂ /-- The `b₄` coefficient of a Weierstrass curve. -/ -@[simp] def b₄ : R := 2 * C.a₄ + C.a₁ * C.a₃ +@[simp] def b₄ : R := 2 * W.a₄ + W.a₁ * W.a₃ /-- The `b₆` coefficient of a Weierstrass curve. -/ -@[simp] def b₆ : R := C.a₃ ^ 2 + 4 * C.a₆ +@[simp] def b₆ : R := W.a₃ ^ 2 + 4 * W.a₆ /-- The `b₈` coefficient of a Weierstrass curve. -/ @[simp] def b₈ : R := -C.a₁ ^ 2 * C.a₆ + 4 * C.a₂ * C.a₆ - C.a₁ * C.a₃ * C.a₄ + C.a₂ * C.a₃ ^ 2 - C.a₄ ^ 2 +W.a₁ ^ 2 * W.a₆ + 4 * W.a₂ * W.a₆ - W.a₁ * W.a₃ * W.a₄ + W.a₂ * W.a₃ ^ 2 - W.a₄ ^ 2 -lemma b_relation : 4 * C.b₈ = C.b₂ * C.b₆ - C.b₄ ^ 2 := by { simp only [b₂, b₄, b₆, b₈], ring1 } +lemma b_relation : 4 * W.b₈ = W.b₂ * W.b₆ - W.b₄ ^ 2 := by { simp only [b₂, b₄, b₆, b₈], ring1 } /-- The `c₄` coefficient of a Weierstrass curve. -/ -@[simp] def c₄ : R := C.b₂ ^ 2 - 24 * C.b₄ +@[simp] def c₄ : R := W.b₂ ^ 2 - 24 * W.b₄ /-- The `c₆` coefficient of a Weierstrass curve. -/ -@[simp] def c₆ : R := -C.b₂ ^ 3 + 36 * C.b₂ * C.b₄ - 216 * C.b₆ +@[simp] def c₆ : R := -W.b₂ ^ 3 + 36 * W.b₂ * W.b₄ - 216 * W.b₆ /-- The discriminant `Δ` of a Weierstrass curve. If `R` is a field, then this polynomial vanishes if and only if the cubic curve cut out by this equation is singular. Sometimes only defined up to sign in the literature; we choose the sign used by the LMFDB. For more discussion, see [the LMFDB page on discriminants](https://www.lmfdb.org/knowledge/show/ec.discriminant). -/ -@[simp] def Δ : R := -C.b₂ ^ 2 * C.b₈ - 8 * C.b₄ ^ 3 - 27 * C.b₆ ^ 2 + 9 * C.b₂ * C.b₄ * C.b₆ +@[simp] def Δ : R := -W.b₂ ^ 2 * W.b₈ - 8 * W.b₄ ^ 3 - 27 * W.b₆ ^ 2 + 9 * W.b₂ * W.b₄ * W.b₆ -lemma c_relation : 1728 * C.Δ = C.c₄ ^ 3 - C.c₆ ^ 2 := +lemma c_relation : 1728 * W.Δ = W.c₄ ^ 3 - W.c₆ ^ 2 := by { simp only [b₂, b₄, b₆, b₈, c₄, c₆, Δ], ring1 } end quantity @@ -124,38 +139,38 @@ section variable_change variables (u : Rˣ) (r s t : R) /-- The Weierstrass curve over `R` induced by an admissible linear change of variables -$(x, y) \mapsto (u^2x + r, u^3y + u^2sx + t)$ for some $u \in R^\times$ and some $r, s, t \in R$. -/ +$(X, Y) \mapsto (u^2X + r, u^3Y + u^2sX + t)$ for some $u \in R^\times$ and some $r, s, t \in R$. -/ @[simps] def variable_change : weierstrass_curve R := -{ a₁ := ↑u⁻¹ * (C.a₁ + 2 * s), - a₂ := ↑u⁻¹ ^ 2 * (C.a₂ - s * C.a₁ + 3 * r - s ^ 2), - a₃ := ↑u⁻¹ ^ 3 * (C.a₃ + r * C.a₁ + 2 * t), - a₄ := ↑u⁻¹ ^ 4 * (C.a₄ - s * C.a₃ + 2 * r * C.a₂ - (t + r * s) * C.a₁ + 3 * r ^ 2 - 2 * s * t), - a₆ := ↑u⁻¹ ^ 6 * (C.a₆ + r * C.a₄ + r ^ 2 * C.a₂ + r ^ 3 - t * C.a₃ - t ^ 2 - r * t * C.a₁) } +{ a₁ := ↑u⁻¹ * (W.a₁ + 2 * s), + a₂ := ↑u⁻¹ ^ 2 * (W.a₂ - s * W.a₁ + 3 * r - s ^ 2), + a₃ := ↑u⁻¹ ^ 3 * (W.a₃ + r * W.a₁ + 2 * t), + a₄ := ↑u⁻¹ ^ 4 * (W.a₄ - s * W.a₃ + 2 * r * W.a₂ - (t + r * s) * W.a₁ + 3 * r ^ 2 - 2 * s * t), + a₆ := ↑u⁻¹ ^ 6 * (W.a₆ + r * W.a₄ + r ^ 2 * W.a₂ + r ^ 3 - t * W.a₃ - t ^ 2 - r * t * W.a₁) } -@[simp] lemma variable_change_b₂ : (C.variable_change u r s t).b₂ = ↑u⁻¹ ^ 2 * (C.b₂ + 12 * r) := +@[simp] lemma variable_change_b₂ : (W.variable_change u r s t).b₂ = ↑u⁻¹ ^ 2 * (W.b₂ + 12 * r) := by { simp only [b₂, variable_change_a₁, variable_change_a₂], ring1 } @[simp] lemma variable_change_b₄ : - (C.variable_change u r s t).b₄ = ↑u⁻¹ ^ 4 * (C.b₄ + r * C.b₂ + 6 * r ^ 2) := + (W.variable_change u r s t).b₄ = ↑u⁻¹ ^ 4 * (W.b₄ + r * W.b₂ + 6 * r ^ 2) := by { simp only [b₂, b₄, variable_change_a₁, variable_change_a₃, variable_change_a₄], ring1 } @[simp] lemma variable_change_b₆ : - (C.variable_change u r s t).b₆ = ↑u⁻¹ ^ 6 * (C.b₆ + 2 * r * C.b₄ + r ^ 2 * C.b₂ + 4 * r ^ 3) := + (W.variable_change u r s t).b₆ = ↑u⁻¹ ^ 6 * (W.b₆ + 2 * r * W.b₄ + r ^ 2 * W.b₂ + 4 * r ^ 3) := by { simp only [b₂, b₄, b₆, variable_change_a₃, variable_change_a₆], ring1 } @[simp] lemma variable_change_b₈ : - (C.variable_change u r s t).b₈ - = ↑u⁻¹ ^ 8 * (C.b₈ + 3 * r * C.b₆ + 3 * r ^ 2 * C.b₄ + r ^ 3 * C.b₂ + 3 * r ^ 4) := + (W.variable_change u r s t).b₈ + = ↑u⁻¹ ^ 8 * (W.b₈ + 3 * r * W.b₆ + 3 * r ^ 2 * W.b₄ + r ^ 3 * W.b₂ + 3 * r ^ 4) := by { simp only [b₂, b₄, b₆, b₈, variable_change_a₁, variable_change_a₂, variable_change_a₃, variable_change_a₄, variable_change_a₆], ring1 } -@[simp] lemma variable_change_c₄ : (C.variable_change u r s t).c₄ = ↑u⁻¹ ^ 4 * C.c₄ := +@[simp] lemma variable_change_c₄ : (W.variable_change u r s t).c₄ = ↑u⁻¹ ^ 4 * W.c₄ := by { simp only [c₄, variable_change_b₂, variable_change_b₄], ring1 } -@[simp] lemma variable_change_c₆ : (C.variable_change u r s t).c₆ = ↑u⁻¹ ^ 6 * C.c₆ := +@[simp] lemma variable_change_c₆ : (W.variable_change u r s t).c₆ = ↑u⁻¹ ^ 6 * W.c₆ := by { simp only [c₆, variable_change_b₂, variable_change_b₄, variable_change_b₆], ring1 } -@[simp] lemma variable_change_Δ : (C.variable_change u r s t).Δ = ↑u⁻¹ ^ 12 * C.Δ := +@[simp] lemma variable_change_Δ : (W.variable_change u r s t).Δ = ↑u⁻¹ ^ 12 * W.Δ := by { dsimp, ring1 } end variable_change @@ -168,29 +183,29 @@ variables (A : Type v) [comm_ring A] [algebra R A] /-- The Weierstrass curve over `R` base changed to `A`. -/ @[simps] def base_change : weierstrass_curve A := -⟨algebra_map R A C.a₁, algebra_map R A C.a₂, algebra_map R A C.a₃, algebra_map R A C.a₄, -algebra_map R A C.a₆⟩ +⟨algebra_map R A W.a₁, algebra_map R A W.a₂, algebra_map R A W.a₃, algebra_map R A W.a₄, +algebra_map R A W.a₆⟩ -@[simp] lemma base_change_b₂ : (C.base_change A).b₂ = algebra_map R A C.b₂ := +@[simp] lemma base_change_b₂ : (W.base_change A).b₂ = algebra_map R A W.b₂ := by { simp only [b₂, base_change_a₁, base_change_a₂], map_simp } -@[simp] lemma base_change_b₄ : (C.base_change A).b₄ = algebra_map R A C.b₄ := +@[simp] lemma base_change_b₄ : (W.base_change A).b₄ = algebra_map R A W.b₄ := by { simp only [b₄, base_change_a₁, base_change_a₃, base_change_a₄], map_simp } -@[simp] lemma base_change_b₆ : (C.base_change A).b₆ = algebra_map R A C.b₆ := +@[simp] lemma base_change_b₆ : (W.base_change A).b₆ = algebra_map R A W.b₆ := by { simp only [b₆, base_change_a₃, base_change_a₆], map_simp } -@[simp] lemma base_change_b₈ : (C.base_change A).b₈ = algebra_map R A C.b₈ := +@[simp] lemma base_change_b₈ : (W.base_change A).b₈ = algebra_map R A W.b₈ := by { simp only [b₈, base_change_a₁, base_change_a₂, base_change_a₃, base_change_a₄, base_change_a₆], map_simp } -@[simp] lemma base_change_c₄ : (C.base_change A).c₄ = algebra_map R A C.c₄ := +@[simp] lemma base_change_c₄ : (W.base_change A).c₄ = algebra_map R A W.c₄ := by { simp only [c₄, base_change_b₂, base_change_b₄], map_simp } -@[simp] lemma base_change_c₆ : (C.base_change A).c₆ = algebra_map R A C.c₆ := +@[simp] lemma base_change_c₆ : (W.base_change A).c₆ = algebra_map R A W.c₆ := by { simp only [c₆, base_change_b₂, base_change_b₄, base_change_b₆], map_simp } -@[simp, nolint simp_nf] lemma base_change_Δ : (C.base_change A).Δ = algebra_map R A C.Δ := +@[simp, nolint simp_nf] lemma base_change_Δ : (W.base_change A).Δ = algebra_map R A W.Δ := by { simp only [Δ, base_change_b₂, base_change_b₄, base_change_b₆, base_change_b₈], map_simp } end base_change @@ -199,23 +214,174 @@ section torsion_polynomial /-! ### 2-torsion polynomials -/ -/-- A cubic polynomial whose discriminant is a multiple of the Weierstrass curve discriminant. -If `C` is an elliptic curve over a field `R` of characteristic different from 2, then its roots over -a splitting field of `R` are precisely the x-coordinates of the non-zero 2-torsion points of `C`. -/ -def two_torsion_polynomial : cubic R := ⟨4, C.b₂, 2 * C.b₄, C.b₆⟩ +/-- A cubic polynomial whose discriminant is a multiple of the Weierstrass curve discriminant. If +`W` is an elliptic curve over a field `R` of characteristic different from 2, then its roots over a +splitting field of `R` are precisely the $X$-coordinates of the non-zero 2-torsion points of `W`. -/ +def two_torsion_polynomial : cubic R := ⟨4, W.b₂, 2 * W.b₄, W.b₆⟩ -lemma two_torsion_polynomial_disc : C.two_torsion_polynomial.disc = 16 * C.Δ := +lemma two_torsion_polynomial_disc : W.two_torsion_polynomial.disc = 16 * W.Δ := by { dsimp [two_torsion_polynomial, cubic.disc], ring1 } lemma two_torsion_polynomial_disc_is_unit [invertible (2 : R)] : - is_unit C.two_torsion_polynomial.disc ↔ is_unit C.Δ := + is_unit W.two_torsion_polynomial.disc ↔ is_unit W.Δ := begin rw [two_torsion_polynomial_disc, is_unit.mul_iff, show (16 : R) = 2 ^ 4, by norm_num1], exact and_iff_right (is_unit_of_invertible $ 2 ^ 4) end +lemma two_torsion_polynomial_disc_ne_zero [nontrivial R] [invertible (2 : R)] (hΔ : is_unit W.Δ) : + W.two_torsion_polynomial.disc ≠ 0 := +(W.two_torsion_polynomial_disc_is_unit.mpr hΔ).ne_zero + end torsion_polynomial +section polynomial + +/-! ### Weierstrass polynomials -/ + +open polynomial + +open_locale polynomial + +/-- The polynomial $W(X, Y) := Y^2 + a_1XY + a_3Y - (X^3 + a_2X^2 + a_4X + a_6)$ associated to a +Weierstrass curve `W` over `R`. For ease of polynomial manipulation, this is represented as a term +of type `R[X][X]`, where the inner variable represents $X$ and the outer variable represents $Y$. -/ +noncomputable def polynomial : R[X][X] := +X ^ 2 + C (C W.a₁ * X + C W.a₃) * X - C (X ^ 3 + C W.a₂ * X ^ 2 + C W.a₄ * X + C W.a₆) + +@[simp] lemma eval_polynomial (x y : R) : + eval x (eval (C y) W.polynomial) + = y ^ 2 + W.a₁ * x * y + W.a₃ * y - (x ^ 3 + W.a₂ * x ^ 2 + W.a₄ * x + W.a₆) := +by { simp only [polynomial], eval_simp, rw [add_mul, ← add_assoc] } + +@[simp] lemma eval_polynomial_zero : eval 0 (eval 0 W.polynomial) = -W.a₆ := +by simp only [← C_0, eval_polynomial, zero_add, zero_sub, mul_zero, zero_pow (nat.zero_lt_succ _)] + +/-- The proposition that an affine point $(x, y)$ lies in `W`. In other words, $W(x, y) = 0$. -/ +def equation (x y : R) : Prop := eval x (eval (C y) W.polynomial) = 0 + +lemma equation_iff' (x y : R) : + W.equation x y ↔ y ^ 2 + W.a₁ * x * y + W.a₃ * y - (x ^ 3 + W.a₂ * x ^ 2 + W.a₄ * x + W.a₆) = 0 := +by rw [equation, eval_polynomial] + +@[simp] lemma equation_iff (x y : R) : + W.equation x y ↔ y ^ 2 + W.a₁ * x * y + W.a₃ * y = x ^ 3 + W.a₂ * x ^ 2 + W.a₄ * x + W.a₆ := +by rw [equation_iff', sub_eq_zero] + +@[simp] lemma equation_zero : W.equation 0 0 ↔ W.a₆ = 0 := +by rw [equation, C_0, eval_polynomial_zero, neg_eq_zero] + +lemma equation_iff_variable_change (x y : R) : + W.equation x y ↔ (W.variable_change 1 x 0 y).equation 0 0 := +begin + rw [equation_iff', ← neg_eq_zero, equation_zero, variable_change_a₆, inv_one, units.coe_one], + congr' 2, + ring1 +end + +/-- The partial derivative $W_X(X, Y)$ of $W(X, Y)$ with respect to $X$. -/ +noncomputable def polynomial_X : R[X][X] := +C (C W.a₁) * X - C (C 3 * X ^ 2 + C (2 * W.a₂) * X + C W.a₄) + +@[simp] lemma eval_polynomial_X (x y : R) : + eval x (eval (C y) W.polynomial_X) = W.a₁ * y - (3 * x ^ 2 + 2 * W.a₂ * x + W.a₄) := +by { simp only [polynomial_X], eval_simp } + +@[simp] lemma eval_polynomial_X_zero : eval 0 (eval 0 W.polynomial_X) = -W.a₄ := +by simp only [← C_0, eval_polynomial_X, zero_add, zero_sub, mul_zero, zero_pow zero_lt_two] + +/-- The partial derivative $W_Y(X, Y)$ of $W(X, Y)$ with respect to $Y$. -/ +noncomputable def polynomial_Y : R[X][X] := C (C 2) * X + C (C W.a₁ * X + C W.a₃) + +@[simp] lemma eval_polynomial_Y (x y : R) : + eval x (eval (C y) W.polynomial_Y) = 2 * y + W.a₁ * x + W.a₃ := +by { simp only [polynomial_Y], eval_simp, rw [← add_assoc] } + +@[simp] lemma eval_polynomial_Y_zero : eval 0 (eval 0 W.polynomial_Y) = W.a₃ := +by simp only [← C_0, eval_polynomial_Y, zero_add, mul_zero] + +/-- The proposition that an affine point $(x, y)$ on `W` is nonsingular. +In other words, either $W_X(x, y) \ne 0$ or $W_Y(x, y) \ne 0$. -/ +def nonsingular (x y : R) : Prop := +eval x (eval (C y) W.polynomial_X) ≠ 0 ∨ eval x (eval (C y) W.polynomial_Y) ≠ 0 + +lemma nonsingular_iff' (x y : R) : + W.nonsingular x y + ↔ W.a₁ * y - (3 * x ^ 2 + 2 * W.a₂ * x + W.a₄) ≠ 0 ∨ 2 * y + W.a₁ * x + W.a₃ ≠ 0 := +by rw [nonsingular, eval_polynomial_X, eval_polynomial_Y] + +@[simp] lemma nonsingular_iff (x y : R) : + W.nonsingular x y ↔ W.a₁ * y ≠ 3 * x ^ 2 + 2 * W.a₂ * x + W.a₄ ∨ y ≠ -y - W.a₁ * x - W.a₃ := +by { rw [nonsingular_iff', sub_ne_zero, ← @sub_ne_zero _ _ y], congr' 3; ring1 } + +@[simp] lemma nonsingular_zero : W.nonsingular 0 0 ↔ W.a₃ ≠ 0 ∨ W.a₄ ≠ 0 := +by rw [nonsingular, C_0, eval_polynomial_X_zero, neg_ne_zero, eval_polynomial_Y_zero, or_comm] + +lemma nonsingular_iff_variable_change (x y : R) : + W.nonsingular x y ↔ (W.variable_change 1 x 0 y).nonsingular 0 0 := +begin + rw [nonsingular_iff', ← neg_ne_zero, or_comm, nonsingular_zero, variable_change_a₃, + variable_change_a₄, inv_one, units.coe_one], + congr' 3, + all_goals { ring1 } +end + +lemma nonsingular_zero_of_Δ_ne_zero (h : W.equation 0 0) (hΔ : W.Δ ≠ 0) : W.nonsingular 0 0 := +by { simp only [equation_zero, nonsingular_zero] at *, contrapose! hΔ, simp [h, hΔ] } + +/-- A Weierstrass curve is nonsingular at every point if its discriminant is non-zero. -/ +lemma nonsingular_of_Δ_ne_zero {x y : R} (h : W.equation x y) (hΔ : W.Δ ≠ 0) : W.nonsingular x y := +(W.nonsingular_iff_variable_change x y).mpr $ + nonsingular_zero_of_Δ_ne_zero _ ((W.equation_iff_variable_change x y).mp h) $ +by rwa [variable_change_Δ, inv_one, units.coe_one, one_pow, one_mul] + +lemma polynomial_eq : W.polynomial = cubic.to_poly + ⟨0, 1, cubic.to_poly ⟨0, 0, W.a₁, W.a₃⟩, cubic.to_poly ⟨-1, -W.a₂, -W.a₄, -W.a₆⟩⟩ := +by { simp only [polynomial, cubic.to_poly, C_0, C_1, C_neg, C_add, C_mul], ring1 } + +lemma polynomial_ne_zero [nontrivial R] : W.polynomial ≠ 0 := +by { rw [polynomial_eq], exact cubic.ne_zero_of_b_ne_zero one_ne_zero } + +lemma polynomial_degree [nontrivial R] : W.polynomial.degree = 2 := +by { rw [polynomial_eq], exact cubic.degree_of_b_ne_zero' one_ne_zero } + +lemma polynomial_nat_degree [nontrivial R] : W.polynomial.nat_degree = 2 := +by { rw [polynomial_eq], exact cubic.nat_degree_of_b_ne_zero' one_ne_zero } + +lemma polynomial_monic : W.polynomial.monic := +by { nontriviality R, simpa only [polynomial_eq] using cubic.monic_of_b_eq_one' } + +lemma polynomial_irreducible [nontrivial R] [no_zero_divisors R] : irreducible W.polynomial := +begin + by_contra h, + rcases (W.polynomial_monic.not_irreducible_iff_exists_add_mul_eq_coeff W.polynomial_nat_degree).mp + h with ⟨f, g, h0, h1⟩, + simp only [polynomial_eq, cubic.coeff_eq_c, cubic.coeff_eq_d] at h0 h1, + apply_fun degree at h0 h1, + rw [cubic.degree_of_a_ne_zero' $ neg_ne_zero.mpr $ one_ne_zero' R, degree_mul] at h0, + apply (h1.symm.le.trans cubic.degree_of_b_eq_zero').not_lt, + rcases nat.with_bot.add_eq_three_iff.mp h0.symm with h | h | h | h, + any_goals { rw [degree_add_eq_left_of_degree_lt]; simp only [h]; dec_trivial }, + any_goals { rw [degree_add_eq_right_of_degree_lt]; simp only [h]; dec_trivial } +end + +/-- The coordinate ring $R[W] := R[X, Y] / \langle W(X, Y) \rangle$ of `W`. -/ +@[reducible] def coordinate_ring : Type u := adjoin_root W.polynomial + +instance [is_domain R] [normalized_gcd_monoid R] : is_domain W.coordinate_ring := +(ideal.quotient.is_domain_iff_prime _).mpr $ +by simpa only [ideal.span_singleton_prime W.polynomial_ne_zero, ← gcd_monoid.irreducible_iff_prime] + using W.polynomial_irreducible + +instance coordinate_ring.is_domain_of_field {F : Type u} [field F] (W : weierstrass_curve F) : + is_domain W.coordinate_ring := +by { classical, apply_instance } + +/-- The function field $R(W) := \mathrm{Frac}(R[W])$ of `W`. -/ +@[reducible] def function_field : Type u := fraction_ring W.coordinate_ring + +end polynomial + end weierstrass_curve /-! ## Elliptic curves -/ @@ -235,9 +401,12 @@ variables [comm_ring R] (E : elliptic_curve R) /-- The j-invariant `j` of an elliptic curve, which is invariant under isomorphisms over `R`. -/ @[simp] def j : R := ↑E.Δ'⁻¹ * E.c₄ ^ 3 -lemma two_torsion_polynomial.disc_ne_zero [nontrivial R] [invertible (2 : R)] : +lemma two_torsion_polynomial_disc_ne_zero [nontrivial R] [invertible (2 : R)] : E.two_torsion_polynomial.disc ≠ 0 := -(E.two_torsion_polynomial_disc_is_unit.mpr $ E.coe_Δ' ▸ E.Δ'.is_unit).ne_zero +E.two_torsion_polynomial_disc_ne_zero $ E.coe_Δ' ▸ E.Δ'.is_unit + +lemma nonsingular [nontrivial R] {x y : R} (h : E.equation x y) : E.nonsingular x y := +E.nonsingular_of_Δ_ne_zero h $ E.coe_Δ' ▸ E.Δ'.ne_zero section variable_change @@ -246,7 +415,7 @@ section variable_change variables (u : Rˣ) (r s t : R) /-- The elliptic curve over `R` induced by an admissible linear change of variables -$(x, y) \mapsto (u^2x + r, u^3y + u^2sx + t)$ for some $u \in R^\times$ and some $r, s, t \in R$. +$(X, Y) \mapsto (u^2X + r, u^3Y + u^2sX + t)$ for some $u \in R^\times$ and some $r, s, t \in R$. When `R` is a field, any two Weierstrass equations isomorphic to `E` are related by this. -/ @[simps] def variable_change : elliptic_curve R := ⟨E.variable_change u r s t, u⁻¹ ^ 12 * E.Δ', @@ -255,8 +424,7 @@ by rw [units.coe_mul, units.coe_pow, coe_Δ', E.variable_change_Δ]⟩ lemma coe_variable_change_Δ' : (↑(E.variable_change u r s t).Δ' : R) = ↑u⁻¹ ^ 12 * E.Δ' := by rw [variable_change_Δ', units.coe_mul, units.coe_pow] -lemma coe_variable_change_Δ'_inv : - (↑(E.variable_change u r s t).Δ'⁻¹ : R) = u ^ 12 * ↑E.Δ'⁻¹ := +lemma coe_variable_change_Δ'_inv : (↑(E.variable_change u r s t).Δ'⁻¹ : R) = u ^ 12 * ↑E.Δ'⁻¹ := by rw [variable_change_Δ', mul_inv, inv_pow, inv_inv, units.coe_mul, units.coe_pow] @[simp] lemma variable_change_j : (E.variable_change u r s t).j = E.j := From d9767b5472a119ce6f34c3007cbcc0cb6a2f2326 Mon Sep 17 00:00:00 2001 From: kkytola Date: Mon, 19 Dec 2022 09:05:54 +0000 Subject: [PATCH 0078/1291] feat(measure_theory/measure/portmanteau): Add lemmas in preparation of implication from Borel condition. (#17962) Add the essential lemmas about the existence of thickenings with null frontier that are needed to prove the portmanteau implication from Borel set condition to closed set condition. --- src/measure_theory/measure/portmanteau.lean | 39 +++++++++++++++++++ .../metric_space/hausdorff_distance.lean | 34 ++++++++-------- 2 files changed, 56 insertions(+), 17 deletions(-) diff --git a/src/measure_theory/measure/portmanteau.lean b/src/measure_theory/measure/portmanteau.lean index 634dd2cf2dbff..7bdf190a5e22d 100644 --- a/src/measure_theory/measure/portmanteau.lean +++ b/src/measure_theory/measure/portmanteau.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kalle Kytölä -/ import measure_theory.measure.probability_measure +import measure_theory.measure.lebesgue /-! # Characterizations of weak convergence of finite measures and probability measures @@ -437,4 +438,42 @@ end end convergence_implies_limsup_closed_le --section +section limit_borel_implies_limsup_closed_le +/-! ### Portmanteau implication: limit condition for Borel sets implies limsup for closed sets + +TODO: The proof of the implication is not yet here. Add it. +-/ + +variables {Ω : Type*} [pseudo_emetric_space Ω] [measurable_space Ω] [opens_measurable_space Ω] + +lemma exists_null_frontier_thickening + (μ : measure Ω) [sigma_finite μ] (s : set Ω) {a b : ℝ} (hab : a < b) : + ∃ r ∈ Ioo a b, μ (frontier (metric.thickening r s)) = 0 := +begin + have mbles : ∀ (r : ℝ), measurable_set (frontier (metric.thickening r s)), + from λ r, (is_closed_frontier).measurable_set, + have disjs := metric.frontier_thickening_disjoint s, + have key := @measure.countable_meas_pos_of_disjoint_Union Ω _ _ μ _ _ mbles disjs, + have aux := @measure_diff_null ℝ _ volume (Ioo a b) _ (set.countable.measure_zero key volume), + have len_pos : 0 < ennreal.of_real (b - a), by simp only [hab, ennreal.of_real_pos, sub_pos], + rw [← real.volume_Ioo, ← aux] at len_pos, + rcases nonempty_of_measure_ne_zero len_pos.ne.symm with ⟨r, ⟨r_in_Ioo, hr⟩⟩, + refine ⟨r, r_in_Ioo, _⟩, + simpa only [mem_set_of_eq, not_lt, le_zero_iff] using hr, +end + +lemma exists_null_frontiers_thickening (μ : measure Ω) [sigma_finite μ] (s : set Ω) : + ∃ (rs : ℕ → ℝ), tendsto rs at_top (𝓝 0) ∧ + ∀ n, 0 < rs n ∧ μ (frontier (metric.thickening (rs n) s)) = 0 := +begin + rcases exists_seq_strict_anti_tendsto (0 : ℝ) with ⟨Rs, ⟨rubbish, ⟨Rs_pos, Rs_lim⟩⟩⟩, + have obs := λ (n : ℕ), exists_null_frontier_thickening μ s (Rs_pos n), + refine ⟨(λ (n : ℕ), (obs n).some), ⟨_, _⟩⟩, + { exact tendsto_of_tendsto_of_tendsto_of_le_of_le tendsto_const_nhds Rs_lim + (λ n, (obs n).some_spec.some.1.le) (λ n, (obs n).some_spec.some.2.le), }, + { exact λ n, ⟨(obs n).some_spec.some.1, (obs n).some_spec.some_spec⟩, }, +end + +end limit_borel_implies_limsup_closed_le --section + end measure_theory --namespace diff --git a/src/topology/metric_space/hausdorff_distance.lean b/src/topology/metric_space/hausdorff_distance.lean index 1643b58e357af..d73ceea17a03b 100644 --- a/src/topology/metric_space/hausdorff_distance.lean +++ b/src/topology/metric_space/hausdorff_distance.lean @@ -883,6 +883,23 @@ lemma mem_thickening_iff_exists_edist_lt {δ : ℝ} (E : set α) (x : α) : x ∈ thickening δ E ↔ ∃ z ∈ E, edist x z < ennreal.of_real δ := inf_edist_lt_iff +/-- The frontier of the (open) thickening of a set is contained in an `inf_edist` level set. -/ +lemma frontier_thickening_subset (E : set α) {δ : ℝ} : + frontier (thickening δ E) ⊆ {x : α | inf_edist x E = ennreal.of_real δ} := +frontier_lt_subset_eq continuous_inf_edist continuous_const + +lemma frontier_thickening_disjoint (A : set α) : + pairwise (disjoint on (λ (r : ℝ), frontier (thickening r A))) := +begin + refine (pairwise_disjoint_on _).2 (λ r₁ r₂ hr, _), + cases le_total r₁ 0 with h₁ h₁, + { simp [thickening_of_nonpos h₁] }, + refine ((disjoint_singleton.2 $ λ h, hr.ne _).preimage _).mono + (frontier_thickening_subset _) (frontier_thickening_subset _), + apply_fun ennreal.to_real at h, + rwa [ennreal.to_real_of_real h₁, ennreal.to_real_of_real (h₁.trans hr.le)] at h +end + variables {X : Type u} [pseudo_metric_space X] /-- A point in a metric space belongs to the (open) `δ`-thickening of a subset `E` if and only if @@ -922,23 +939,6 @@ begin ... ≤ R + δ : add_le_add (hR zE) hz.le end -/-- The frontier of the (open) thickening of a set is contained in an `inf_edist` level set. -/ -lemma frontier_thickening_subset (E : set α) {δ : ℝ} : - frontier (thickening δ E) ⊆ {x : α | inf_edist x E = ennreal.of_real δ} := -frontier_lt_subset_eq continuous_inf_edist continuous_const - -lemma frontier_thickening_disjoint (A : set X) : - pairwise (disjoint on (λ (r : ℝ), frontier (thickening r A))) := -begin - refine (pairwise_disjoint_on _).2 (λ r₁ r₂ hr, _), - cases le_total r₁ 0 with h₁ h₁, - { simp [thickening_of_nonpos h₁] }, - refine ((disjoint_singleton.2 $ λ h, hr.ne _).preimage _).mono - (frontier_thickening_subset _) (frontier_thickening_subset _), - apply_fun ennreal.to_real at h, - rwa [ennreal.to_real_of_real h₁, ennreal.to_real_of_real (h₁.trans hr.le)] at h -end - end thickening --section section cthickening From 550b58538991c8977703fdeb7c9d51a5aa27df11 Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Mon, 19 Dec 2022 16:43:24 +0000 Subject: [PATCH 0079/1291] chore(algebraic_geometry/elliptic_curve/weierstrass): add disclaimer for coordinate_ring irreducibility (#17977) Also rename `coe_inv` lemmas to be consistent with those generated by `@[simps]`. Co-authored-by: Anne Baanen --- .../elliptic_curve/weierstrass.lean | 24 +++++++++++++------ 1 file changed, 17 insertions(+), 7 deletions(-) diff --git a/src/algebraic_geometry/elliptic_curve/weierstrass.lean b/src/algebraic_geometry/elliptic_curve/weierstrass.lean index 78a0cec827478..ac7082cf315e8 100644 --- a/src/algebraic_geometry/elliptic_curve/weierstrass.lean +++ b/src/algebraic_geometry/elliptic_curve/weierstrass.lean @@ -365,13 +365,23 @@ begin any_goals { rw [degree_add_eq_right_of_degree_lt]; simp only [h]; dec_trivial } end -/-- The coordinate ring $R[W] := R[X, Y] / \langle W(X, Y) \rangle$ of `W`. -/ -@[reducible] def coordinate_ring : Type u := adjoin_root W.polynomial +/-- The coordinate ring $R[W] := R[X, Y] / \langle W(X, Y) \rangle$ of `W`. + +Note that `derive comm_ring` generates a reducible instance of `comm_ring` for `coordinate_ring`. +In certain circumstances this might be extremely slow, because all instances in its definition are +unified exponentially many times. In this case, one solution is to manually add the local attribute +`local attribute [irreducible] coordinate_ring.comm_ring` to block this type-level unification. + +TODO Lean 4: verify if the new def-eq cache (lean4#1102) fixed this issue. + +See Zulip thread: +https://leanprover.zulipchat.com/#narrow/stream/116395-maths/topic/.E2.9C.94.20class_group.2Emk -/ +@[derive [inhabited, comm_ring]] def coordinate_ring : Type u := adjoin_root W.polynomial instance [is_domain R] [normalized_gcd_monoid R] : is_domain W.coordinate_ring := (ideal.quotient.is_domain_iff_prime _).mpr $ by simpa only [ideal.span_singleton_prime W.polynomial_ne_zero, ← gcd_monoid.irreducible_iff_prime] - using W.polynomial_irreducible + using W.polynomial_irreducible instance coordinate_ring.is_domain_of_field {F : Type u} [field F] (W : weierstrass_curve F) : is_domain W.coordinate_ring := @@ -424,12 +434,12 @@ by rw [units.coe_mul, units.coe_pow, coe_Δ', E.variable_change_Δ]⟩ lemma coe_variable_change_Δ' : (↑(E.variable_change u r s t).Δ' : R) = ↑u⁻¹ ^ 12 * E.Δ' := by rw [variable_change_Δ', units.coe_mul, units.coe_pow] -lemma coe_variable_change_Δ'_inv : (↑(E.variable_change u r s t).Δ'⁻¹ : R) = u ^ 12 * ↑E.Δ'⁻¹ := +lemma coe_inv_variable_change_Δ' : (↑(E.variable_change u r s t).Δ'⁻¹ : R) = u ^ 12 * ↑E.Δ'⁻¹ := by rw [variable_change_Δ', mul_inv, inv_pow, inv_inv, units.coe_mul, units.coe_pow] @[simp] lemma variable_change_j : (E.variable_change u r s t).j = E.j := begin - rw [j, coe_variable_change_Δ'_inv], + rw [j, coe_inv_variable_change_Δ'], have hu : (u * ↑u⁻¹ : R) ^ 12 = 1 := by rw [u.mul_inv, one_pow], linear_combination E.j * hu with { normalization_tactic := `[dsimp, ring1] } end @@ -449,10 +459,10 @@ by rw [units.coe_map, ring_hom.coe_monoid_hom, coe_Δ', E.base_change_Δ]⟩ lemma coe_base_change_Δ' : ↑(E.base_change A).Δ' = algebra_map R A E.Δ' := rfl -lemma coe_base_change_Δ'_inv : ↑(E.base_change A).Δ'⁻¹ = algebra_map R A ↑E.Δ'⁻¹ := rfl +lemma coe_inv_base_change_Δ' : ↑(E.base_change A).Δ'⁻¹ = algebra_map R A ↑E.Δ'⁻¹ := rfl @[simp] lemma base_change_j : (E.base_change A).j = algebra_map R A E.j := -by { simp only [j, coe_base_change_Δ'_inv, base_change_to_weierstrass_curve, E.base_change_c₄], +by { simp only [j, coe_inv_base_change_Δ', base_change_to_weierstrass_curve, E.base_change_c₄], map_simp } end base_change From a7aa4a2e6c4ebadf4974167c7c578772b60431c4 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 20 Dec 2022 09:02:03 +0000 Subject: [PATCH 0080/1291] fix(scripts/port_comments): handle moved files (#17978) --- scripts/add_port_comments.py | 22 +++++++++++++++++----- 1 file changed, 17 insertions(+), 5 deletions(-) diff --git a/scripts/add_port_comments.py b/scripts/add_port_comments.py index 7bf81c9db5742..35ef63c89ccb7 100644 --- a/scripts/add_port_comments.py +++ b/scripts/add_port_comments.py @@ -92,15 +92,20 @@ def fname_for(import_path: str) -> Path: missing_docstrings = [] +missing_files = [] for iname, f_status in status.file_statuses.items(): if f_status.ported: fname = fname_for(iname) - with open(fname) as f: - fcontent = f.read() + try: + with open(fname) as f: + fcontent = f.read() + except FileNotFoundError: + missing_files.append((iname, fname)) + continue try: new_fcontent = add_port_status(fcontent, f_status) except NoModuleDocstringError: - missing_docstrings.append(fname) + missing_docstrings.append((iname, fname)) continue if new_fcontent == fcontent: continue @@ -110,6 +115,13 @@ def fname_for(import_path: str) -> Path: if missing_docstrings: print('\n---') print('The following files have no module docstring, so I have not added a message in this PR') - for fname in missing_docstrings: - print(f'* [`{fname}`](https://github.com/leanprover-community/mathlib/blob/master/{fname})') + for iname, fname in missing_docstrings: + print(f'* [`{iname}`](https://github.com/leanprover-community/mathlib/blob/master/{fname})') print('\nPlease make a PR to add a module docstring (for Lean3 and Lean4!), then I will add the freeze comment next time.') +if missing_files: + print('\n---') + print('The following files no longer exist in Lean 3\' mathlib, so I have not added a message in this PR') + for iname, fname in missing_files: + f_status = status.file_statuses[iname] + print(f'* [`{iname}`](https://github.com/leanprover-community/mathlib/blob/{f_status.mathlib3_hash}/{fname})') + print('\nIn future we should find where they moved to, and check that the files are still in sync.') From 09fa0b0cad2a2622ded89be5f9dee80b8e8296d8 Mon Sep 17 00:00:00 2001 From: Felix-Weilacher Date: Tue, 20 Dec 2022 17:39:22 +0000 Subject: [PATCH 0081/1291] refactor(measure_theory/measurable_space): utilize measurable_embedding more (#17980) Change the hypothesis in `measurable_equiv.set.image` and `measurable_equiv.set.range` to use `measurable_embedding`. Also remove measurable_embedding.equiv_range which is redundant. --- src/measure_theory/measurable_space.lean | 45 +++++++++--------------- 1 file changed, 17 insertions(+), 28 deletions(-) diff --git a/src/measure_theory/measurable_space.lean b/src/measure_theory/measurable_space.lean index fd124a1b8b473..790360725b97b 100644 --- a/src/measure_theory/measurable_space.lean +++ b/src/measure_theory/measurable_space.lean @@ -1112,27 +1112,6 @@ def set.singleton (a : α) : ({a} : set α) ≃ᵐ unit := measurable_to_fun := measurable_const, measurable_inv_fun := measurable_const } -/-- A set is equivalent to its image under a function `f` as measurable spaces, - if `f` is an injective measurable function that sends measurable sets to measurable sets. -/ -noncomputable def set.image (f : α → β) (s : set α) (hf : injective f) - (hfm : measurable f) (hfi : ∀ s, measurable_set s → measurable_set (f '' s)) : s ≃ᵐ (f '' s) := -{ to_equiv := equiv.set.image f s hf, - measurable_to_fun := (hfm.comp measurable_id.subtype_coe).subtype_mk, - measurable_inv_fun := - begin - rintro t ⟨u, hu, rfl⟩, simp [preimage_preimage, set.image_symm_preimage hf], - exact measurable_subtype_coe (hfi u hu) - end } - -/-- The domain of `f` is equivalent to its range as measurable spaces, - if `f` is an injective measurable function that sends measurable sets to measurable sets. -/ -noncomputable def set.range (f : α → β) (hf : injective f) (hfm : measurable f) - (hfi : ∀ s, measurable_set s → measurable_set (f '' s)) : - α ≃ᵐ (range f) := -(measurable_equiv.set.univ _).symm.trans $ - (measurable_equiv.set.image f univ hf hfm hfi).trans $ - measurable_equiv.cast (by rw image_univ) (by rw image_univ) - /-- `α` is equivalent to its image in `α ⊕ β` as measurable spaces. -/ def set.range_inl : (range sum.inl : set (α ⊕ β)) ≃ᵐ α := { to_fun := λ ab, match ab with @@ -1275,14 +1254,24 @@ namespace measurable_embedding variables [measurable_space α] [measurable_space β] [measurable_space γ] {f : α → β} -/-- A measurable embedding defines a measurable equivalence between its domain -and its range. -/ -noncomputable def equiv_range (f : α → β) (hf : measurable_embedding f) : - α ≃ᵐ range f := -{ to_equiv := equiv.of_injective f hf.injective, - measurable_to_fun := hf.measurable.subtype_mk, +/-- A set is equivalent to its image under a function `f` as measurable spaces, + if `f` is a measurable embedding -/ +noncomputable def equiv_image (s : set α) (hf : measurable_embedding f) : + s ≃ᵐ (f '' s) := +{ to_equiv := equiv.set.image f s hf.injective, + measurable_to_fun := (hf.measurable.comp measurable_id.subtype_coe).subtype_mk, measurable_inv_fun := - by { rw coe_of_injective_symm, exact hf.measurable_range_splitting } } + begin + rintro t ⟨u, hu, rfl⟩, simp [preimage_preimage, set.image_symm_preimage hf.injective], + exact measurable_subtype_coe (hf.measurable_set_image' hu) + end } + +/-- The domain of `f` is equivalent to its range as measurable spaces, + if `f` is a measurable embedding -/ +noncomputable def equiv_range (hf : measurable_embedding f) : α ≃ᵐ (range f) := +(measurable_equiv.set.univ _).symm.trans $ + (hf.equiv_image univ).trans $ + measurable_equiv.cast (by rw image_univ) (by rw image_univ) lemma of_measurable_inverse_on_range {g : range f → α} (hf₁ : measurable f) (hf₂ : measurable_set (range f)) (hg : measurable g) From ba2245edf0c8bb155f1569fd9b9492a9b384cde6 Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Tue, 20 Dec 2022 19:05:52 +0000 Subject: [PATCH 0082/1291] feat(analysis/convex/between): betweenness and affine combinations (#17912) Add some lemmas about betweenness for affine combinations in relation to the coefficients in those combinations. --- src/analysis/convex/between.lean | 110 +++++++++++++++++++++++++++++++ 1 file changed, 110 insertions(+) diff --git a/src/analysis/convex/between.lean b/src/analysis/convex/between.lean index 3e0f9fc5e657a..11d1e62efd2cc 100644 --- a/src/analysis/convex/between.lean +++ b/src/analysis/convex/between.lean @@ -25,6 +25,8 @@ variables (R : Type*) {V V' P P' : Type*} open affine_equiv affine_map +open_locale big_operators + section ordered_ring variables [ordered_ring R] [add_comm_group V] [module R V] [add_torsor V P] @@ -452,6 +454,39 @@ lemma sbtw.trans_wbtw_right_ne [no_zero_smul_divisors R V] {w x y z : P} (h₁ : (h₂ : wbtw R x y z) : w ≠ y := h₁.wbtw.trans_right_ne h₂ h₁.left_ne +/- Calls to `affine_combination` are slow to elaborate (generally, not just for this lemma), and +without the use of `@finset.affine_combination R V _ _ _ _ _ _` for at least three of the six +calls in this lemma statement, elaboration of the statement times out (even if the proof is +replaced by `sorry`). -/ +lemma sbtw.affine_combination_of_mem_affine_span_pair [no_zero_divisors R] + [no_zero_smul_divisors R V] {ι : Type*} {p : ι → P} (ha : affine_independent R p) + {w w₁ w₂ : ι → R} {s : finset ι} (hw : ∑ i in s, w i = 1) (hw₁ : ∑ i in s, w₁ i = 1) + (hw₂ : ∑ i in s, w₂ i = 1) + (h : s.affine_combination p w ∈ + line[R, s.affine_combination p w₁, s.affine_combination p w₂]) {i : ι} (his : i ∈ s) + (hs : sbtw R (w₁ i) (w i) (w₂ i)) : + sbtw R (@finset.affine_combination R V _ _ _ _ _ _ s p w₁) + (@finset.affine_combination R V _ _ _ _ _ _ s p w) + (@finset.affine_combination R V _ _ _ _ _ _ s p w₂) := +begin + rw affine_combination_mem_affine_span_pair ha hw hw₁ hw₂ at h, + rcases h with ⟨r, hr⟩, + dsimp only at hr, + rw [hr i his, sbtw_mul_sub_add_iff] at hs, + change ∀ i ∈ s, w i = ((r • (w₂ - w₁) + w₁) i) at hr, + rw s.affine_combination_congr hr (λ _ _, rfl), + dsimp only, + rw [←s.weighted_vsub_vadd_affine_combination, s.weighted_vsub_const_smul, + ←s.affine_combination_vsub, ←line_map_apply, sbtw_line_map_iff, and_iff_left hs.2, + ←@vsub_ne_zero V, s.affine_combination_vsub], + intro hz, + have hw₁w₂ : ∑ i in s, (w₁ - w₂) i = 0, + { simp_rw [pi.sub_apply, finset.sum_sub_distrib, hw₁, hw₂, sub_self] }, + refine hs.1 _, + have ha' := ha s (w₁ - w₂) hw₁w₂ hz i his, + rwa [pi.sub_apply, sub_eq_zero] at ha' +end + end ordered_ring section strict_ordered_comm_ring @@ -487,6 +522,81 @@ end end strict_ordered_comm_ring +section linear_ordered_ring + +variables [linear_ordered_ring R] [add_comm_group V] [module R V] [add_torsor V P] + +include V + +variables {R} + +/-- Suppose lines from two vertices of a triangle to interior points of the opposite side meet at +`p`. Then `p` lies in the interior of the first (and by symmetry the other) segment from a +vertex to the point on the opposite side. -/ +lemma sbtw_of_sbtw_of_sbtw_of_mem_affine_span_pair [no_zero_smul_divisors R V] + {t : affine.triangle R P} {i₁ i₂ i₃ : fin 3} (h₁₂ : i₁ ≠ i₂) {p₁ p₂ p : P} + (h₁ : sbtw R (t.points i₂) p₁ (t.points i₃)) (h₂ : sbtw R (t.points i₁) p₂ (t.points i₃)) + (h₁' : p ∈ line[R, t.points i₁, p₁]) (h₂' : p ∈ line[R, t.points i₂, p₂]) : + sbtw R (t.points i₁) p p₁ := +begin + -- Should not be needed; see comments on local instances in `data.sign`. + letI : decidable_rel ((<) : R → R → Prop) := linear_ordered_ring.decidable_lt, + have h₁₃ : i₁ ≠ i₃, { rintro rfl, simpa using h₂ }, + have h₂₃ : i₂ ≠ i₃, { rintro rfl, simpa using h₁ }, + have h3 : ∀ i : fin 3, i = i₁ ∨ i = i₂ ∨ i = i₃, { clear h₁ h₂ h₁' h₂', dec_trivial! }, + have hu : (finset.univ : finset (fin 3)) = {i₁, i₂, i₃}, { clear h₁ h₂ h₁' h₂', dec_trivial! }, + have hp : p ∈ affine_span R (set.range t.points), + { have hle : line[R, t.points i₁, p₁] ≤ affine_span R (set.range t.points), + { refine affine_span_pair_le_of_mem_of_mem (mem_affine_span _ (set.mem_range_self _)) _, + have hle : line[R, t.points i₂, t.points i₃] ≤ affine_span R (set.range t.points), + { refine affine_span_mono _ _, simp [set.insert_subset] }, + rw affine_subspace.le_def' at hle, + exact hle _ h₁.wbtw.mem_affine_span }, + rw affine_subspace.le_def' at hle, + exact hle _ h₁' }, + have h₁i := h₁.mem_image_Ioo, + have h₂i := h₂.mem_image_Ioo, + rw set.mem_image at h₁i h₂i, + rcases h₁i with ⟨r₁, ⟨hr₁0, hr₁1⟩, rfl⟩, + rcases h₂i with ⟨r₂, ⟨hr₂0, hr₂1⟩, rfl⟩, + rcases eq_affine_combination_of_mem_affine_span_of_fintype hp with ⟨w, hw, rfl⟩, + have h₁s := sign_eq_of_affine_combination_mem_affine_span_single_line_map t.independent hw + (finset.mem_univ _) (finset.mem_univ _) (finset.mem_univ _) h₁₂ h₁₃ h₂₃ hr₁0 hr₁1 h₁', + have h₂s := sign_eq_of_affine_combination_mem_affine_span_single_line_map t.independent hw + (finset.mem_univ _) (finset.mem_univ _) (finset.mem_univ _) h₁₂.symm h₂₃ h₁₃ hr₂0 hr₂1 h₂', + dsimp only at h₁s h₂s, + rw [←finset.univ.affine_combination_affine_combination_single_weights R t.points + (finset.mem_univ i₁), + ←finset.univ.affine_combination_affine_combination_line_map_weights t.points + (finset.mem_univ _) (finset.mem_univ _)] at ⊢ h₁', + refine sbtw.affine_combination_of_mem_affine_span_pair t.independent hw + (finset.univ.sum_affine_combination_single_weights R (finset.mem_univ _)) + (finset.univ.sum_affine_combination_line_map_weights (finset.mem_univ _) (finset.mem_univ _) _) + h₁' (finset.mem_univ i₁) _, + rw [finset.affine_combination_single_weights_apply_self, + finset.affine_combination_line_map_weights_apply_of_ne h₁₂ h₁₃, sbtw_one_zero_iff], + have hs : ∀ i : fin 3, sign (w i) = sign (w i₃), + { intro i, + rcases h3 i with rfl | rfl | rfl, + { exact h₂s }, + { exact h₁s }, + { refl } }, + have hss : sign (∑ i, w i) = 1, { simp [hw] }, + have hs' := sign_sum (finset.univ_nonempty) (sign (w i₃)) (λ i _, hs i), + rw hs' at hss, + simp_rw [hss, sign_eq_one_iff] at hs, + refine ⟨hs i₁, _⟩, + rw hu at hw, + rw [finset.sum_insert, finset.sum_insert, finset.sum_singleton] at hw, + { by_contra hle, + rw not_lt at hle, + exact (hle.trans_lt (lt_add_of_pos_right _ (left.add_pos (hs i₂) (hs i₃)))).ne' hw }, + { simp [h₂₃] }, + { simp [h₁₂, h₁₃] } +end + +end linear_ordered_ring + section linear_ordered_field variables [linear_ordered_field R] [add_comm_group V] [module R V] [add_torsor V P] From e8a452533ab40cd726031244874c482519a4b187 Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Wed, 21 Dec 2022 05:40:50 +0000 Subject: [PATCH 0083/1291] feat(geometry/euclidean/angle/oriented/right_angle): trigonometric lemmas (#17683) Add lemmas about trigonometric functions in relation to right-angled triangles with oriented angles. These are generally analogous to the unoriented case, but note that the property of an angle being a right angle is not symmetric with respect to reversing the order of the vectors / points in the oriented case (a positive right angle becomes a negative right angle). Thus, one unoriented lemma tends to correspond to two lemmas in the oriented case, in order to state the trigonometric results for both angles other than the right angle, since those results are no longer so trivially equivalent. --- .../euclidean/angle/oriented/right_angle.lean | 878 ++++++++++++++++++ 1 file changed, 878 insertions(+) create mode 100644 src/geometry/euclidean/angle/oriented/right_angle.lean diff --git a/src/geometry/euclidean/angle/oriented/right_angle.lean b/src/geometry/euclidean/angle/oriented/right_angle.lean new file mode 100644 index 0000000000000..52462e99f46d8 --- /dev/null +++ b/src/geometry/euclidean/angle/oriented/right_angle.lean @@ -0,0 +1,878 @@ +/- +Copyright (c) 2022 Joseph Myers. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joseph Myers +-/ +import geometry.euclidean.angle.oriented.affine +import geometry.euclidean.angle.unoriented.right_angle + +/-! +# Oriented angles in right-angled triangles. + +This file proves basic geometrical results about distances and oriented angles in (possibly +degenerate) right-angled triangles in real inner product spaces and Euclidean affine spaces. + +-/ + +noncomputable theory + +open_locale euclidean_geometry +open_locale real +open_locale real_inner_product_space + +namespace orientation + +open finite_dimensional + +variables {V : Type*} [inner_product_space ℝ V] +variables [hd2 : fact (finrank ℝ V = 2)] (o : orientation ℝ V (fin 2)) +include hd2 o + +/-- An angle in a right-angled triangle expressed using `arccos`. -/ +lemma oangle_add_right_eq_arccos_of_oangle_eq_pi_div_two {x y : V} (h : o.oangle x y = ↑(π / 2)) : + o.oangle x (x + y) = real.arccos (‖x‖ / ‖x + y‖) := +begin + have hs : (o.oangle x (x + y)).sign = 1, + { rw [oangle_sign_add_right, h, real.angle.sign_coe_pi_div_two] }, + rw [o.oangle_eq_angle_of_sign_eq_one hs, + inner_product_geometry.angle_add_eq_arccos_of_inner_eq_zero + (o.inner_eq_zero_of_oangle_eq_pi_div_two h)] +end + +/-- An angle in a right-angled triangle expressed using `arccos`. -/ +lemma oangle_add_left_eq_arccos_of_oangle_eq_pi_div_two {x y : V} (h : o.oangle x y = ↑(π / 2)) : + o.oangle (x + y) y = real.arccos (‖y‖ / ‖x + y‖) := +begin + rw [←neg_inj, oangle_rev, ←oangle_neg_orientation_eq_neg, neg_inj] at h ⊢, + rw add_comm, + exact (-o).oangle_add_right_eq_arccos_of_oangle_eq_pi_div_two h +end + +/-- An angle in a right-angled triangle expressed using `arcsin`. -/ +lemma oangle_add_right_eq_arcsin_of_oangle_eq_pi_div_two {x y : V} (h : o.oangle x y = ↑(π / 2)) : + o.oangle x (x + y) = real.arcsin (‖y‖ / ‖x + y‖) := +begin + have hs : (o.oangle x (x + y)).sign = 1, + { rw [oangle_sign_add_right, h, real.angle.sign_coe_pi_div_two] }, + rw [o.oangle_eq_angle_of_sign_eq_one hs, + inner_product_geometry.angle_add_eq_arcsin_of_inner_eq_zero + (o.inner_eq_zero_of_oangle_eq_pi_div_two h) + (or.inl (o.left_ne_zero_of_oangle_eq_pi_div_two h))] +end + +/-- An angle in a right-angled triangle expressed using `arcsin`. -/ +lemma oangle_add_left_eq_arcsin_of_oangle_eq_pi_div_two {x y : V} (h : o.oangle x y = ↑(π / 2)) : + o.oangle (x + y) y = real.arcsin (‖x‖ / ‖x + y‖) := +begin + rw [←neg_inj, oangle_rev, ←oangle_neg_orientation_eq_neg, neg_inj] at h ⊢, + rw add_comm, + exact (-o).oangle_add_right_eq_arcsin_of_oangle_eq_pi_div_two h +end + +/-- An angle in a right-angled triangle expressed using `arctan`. -/ +lemma oangle_add_right_eq_arctan_of_oangle_eq_pi_div_two {x y : V} (h : o.oangle x y = ↑(π / 2)) : + o.oangle x (x + y) = real.arctan (‖y‖ / ‖x‖) := +begin + have hs : (o.oangle x (x + y)).sign = 1, + { rw [oangle_sign_add_right, h, real.angle.sign_coe_pi_div_two] }, + rw [o.oangle_eq_angle_of_sign_eq_one hs, + inner_product_geometry.angle_add_eq_arctan_of_inner_eq_zero + (o.inner_eq_zero_of_oangle_eq_pi_div_two h) (o.left_ne_zero_of_oangle_eq_pi_div_two h)] +end + +/-- An angle in a right-angled triangle expressed using `arctan`. -/ +lemma oangle_add_left_eq_arctan_of_oangle_eq_pi_div_two {x y : V} (h : o.oangle x y = ↑(π / 2)) : + o.oangle (x + y) y = real.arctan (‖x‖ / ‖y‖) := +begin + rw [←neg_inj, oangle_rev, ←oangle_neg_orientation_eq_neg, neg_inj] at h ⊢, + rw add_comm, + exact (-o).oangle_add_right_eq_arctan_of_oangle_eq_pi_div_two h +end + +/-- The cosine of an angle in a right-angled triangle as a ratio of sides. -/ +lemma cos_oangle_add_right_of_oangle_eq_pi_div_two {x y : V} (h : o.oangle x y = ↑(π / 2)) : + real.angle.cos (o.oangle x (x + y)) = ‖x‖ / ‖x + y‖ := +begin + have hs : (o.oangle x (x + y)).sign = 1, + { rw [oangle_sign_add_right, h, real.angle.sign_coe_pi_div_two] }, + rw [o.oangle_eq_angle_of_sign_eq_one hs, real.angle.cos_coe, + inner_product_geometry.cos_angle_add_of_inner_eq_zero + (o.inner_eq_zero_of_oangle_eq_pi_div_two h)] +end + +/-- The cosine of an angle in a right-angled triangle as a ratio of sides. -/ +lemma cos_oangle_add_left_of_oangle_eq_pi_div_two {x y : V} (h : o.oangle x y = ↑(π / 2)) : + real.angle.cos (o.oangle (x + y) y) = ‖y‖ / ‖x + y‖ := +begin + rw [←neg_inj, oangle_rev, ←oangle_neg_orientation_eq_neg, neg_inj] at h ⊢, + rw add_comm, + exact (-o).cos_oangle_add_right_of_oangle_eq_pi_div_two h +end + +/-- The sine of an angle in a right-angled triangle as a ratio of sides. -/ +lemma sin_oangle_add_right_of_oangle_eq_pi_div_two {x y : V} (h : o.oangle x y = ↑(π / 2)) : + real.angle.sin (o.oangle x (x + y)) = ‖y‖ / ‖x + y‖ := +begin + have hs : (o.oangle x (x + y)).sign = 1, + { rw [oangle_sign_add_right, h, real.angle.sign_coe_pi_div_two] }, + rw [o.oangle_eq_angle_of_sign_eq_one hs, real.angle.sin_coe, + inner_product_geometry.sin_angle_add_of_inner_eq_zero + (o.inner_eq_zero_of_oangle_eq_pi_div_two h) + (or.inl (o.left_ne_zero_of_oangle_eq_pi_div_two h))] +end + +/-- The sine of an angle in a right-angled triangle as a ratio of sides. -/ +lemma sin_oangle_add_left_of_oangle_eq_pi_div_two {x y : V} (h : o.oangle x y = ↑(π / 2)) : + real.angle.sin (o.oangle (x + y) y) = ‖x‖ / ‖x + y‖ := +begin + rw [←neg_inj, oangle_rev, ←oangle_neg_orientation_eq_neg, neg_inj] at h ⊢, + rw add_comm, + exact (-o).sin_oangle_add_right_of_oangle_eq_pi_div_two h +end + +/-- The tangent of an angle in a right-angled triangle as a ratio of sides. -/ +lemma tan_oangle_add_right_of_oangle_eq_pi_div_two {x y : V} (h : o.oangle x y = ↑(π / 2)) : + real.angle.tan (o.oangle x (x + y)) = ‖y‖ / ‖x‖ := +begin + have hs : (o.oangle x (x + y)).sign = 1, + { rw [oangle_sign_add_right, h, real.angle.sign_coe_pi_div_two] }, + rw [o.oangle_eq_angle_of_sign_eq_one hs, real.angle.tan_coe, + inner_product_geometry.tan_angle_add_of_inner_eq_zero + (o.inner_eq_zero_of_oangle_eq_pi_div_two h)] +end + +/-- The tangent of an angle in a right-angled triangle as a ratio of sides. -/ +lemma tan_oangle_add_left_of_oangle_eq_pi_div_two {x y : V} (h : o.oangle x y = ↑(π / 2)) : + real.angle.tan (o.oangle (x + y) y) = ‖x‖ / ‖y‖ := +begin + rw [←neg_inj, oangle_rev, ←oangle_neg_orientation_eq_neg, neg_inj] at h ⊢, + rw add_comm, + exact (-o).tan_oangle_add_right_of_oangle_eq_pi_div_two h +end + +/-- The cosine of an angle in a right-angled triangle multiplied by the hypotenuse equals the +adjacent side. -/ +lemma cos_oangle_add_right_mul_norm_of_oangle_eq_pi_div_two {x y : V} + (h : o.oangle x y = ↑(π / 2)) : real.angle.cos (o.oangle x (x + y)) * ‖x + y‖ = ‖x‖ := +begin + have hs : (o.oangle x (x + y)).sign = 1, + { rw [oangle_sign_add_right, h, real.angle.sign_coe_pi_div_two] }, + rw [o.oangle_eq_angle_of_sign_eq_one hs, real.angle.cos_coe, + inner_product_geometry.cos_angle_add_mul_norm_of_inner_eq_zero + (o.inner_eq_zero_of_oangle_eq_pi_div_two h)] +end + +/-- The cosine of an angle in a right-angled triangle multiplied by the hypotenuse equals the +adjacent side. -/ +lemma cos_oangle_add_left_mul_norm_of_oangle_eq_pi_div_two {x y : V} + (h : o.oangle x y = ↑(π / 2)) : real.angle.cos (o.oangle (x + y) y) * ‖x + y‖ = ‖y‖ := +begin + rw [←neg_inj, oangle_rev, ←oangle_neg_orientation_eq_neg, neg_inj] at h ⊢, + rw add_comm, + exact (-o).cos_oangle_add_right_mul_norm_of_oangle_eq_pi_div_two h +end + +/-- The sine of an angle in a right-angled triangle multiplied by the hypotenuse equals the +opposite side. -/ +lemma sin_oangle_add_right_mul_norm_of_oangle_eq_pi_div_two {x y : V} + (h : o.oangle x y = ↑(π / 2)) : real.angle.sin (o.oangle x (x + y)) * ‖x + y‖ = ‖y‖ := +begin + have hs : (o.oangle x (x + y)).sign = 1, + { rw [oangle_sign_add_right, h, real.angle.sign_coe_pi_div_two] }, + rw [o.oangle_eq_angle_of_sign_eq_one hs, real.angle.sin_coe, + inner_product_geometry.sin_angle_add_mul_norm_of_inner_eq_zero + (o.inner_eq_zero_of_oangle_eq_pi_div_two h)] +end + +/-- The sine of an angle in a right-angled triangle multiplied by the hypotenuse equals the +opposite side. -/ +lemma sin_oangle_add_left_mul_norm_of_oangle_eq_pi_div_two {x y : V} + (h : o.oangle x y = ↑(π / 2)) : real.angle.sin (o.oangle (x + y) y) * ‖x + y‖ = ‖x‖ := +begin + rw [←neg_inj, oangle_rev, ←oangle_neg_orientation_eq_neg, neg_inj] at h ⊢, + rw add_comm, + exact (-o).sin_oangle_add_right_mul_norm_of_oangle_eq_pi_div_two h +end + +/-- The tangent of an angle in a right-angled triangle multiplied by the adjacent side equals +the opposite side. -/ +lemma tan_oangle_add_right_mul_norm_of_oangle_eq_pi_div_two {x y : V} + (h : o.oangle x y = ↑(π / 2)) : real.angle.tan (o.oangle x (x + y)) * ‖x‖ = ‖y‖ := +begin + have hs : (o.oangle x (x + y)).sign = 1, + { rw [oangle_sign_add_right, h, real.angle.sign_coe_pi_div_two] }, + rw [o.oangle_eq_angle_of_sign_eq_one hs, real.angle.tan_coe, + inner_product_geometry.tan_angle_add_mul_norm_of_inner_eq_zero + (o.inner_eq_zero_of_oangle_eq_pi_div_two h) + (or.inl (o.left_ne_zero_of_oangle_eq_pi_div_two h))] +end + +/-- The tangent of an angle in a right-angled triangle multiplied by the adjacent side equals +the opposite side. -/ +lemma tan_oangle_add_left_mul_norm_of_oangle_eq_pi_div_two {x y : V} + (h : o.oangle x y = ↑(π / 2)) : real.angle.tan (o.oangle (x + y) y) * ‖y‖ = ‖x‖ := +begin + rw [←neg_inj, oangle_rev, ←oangle_neg_orientation_eq_neg, neg_inj] at h ⊢, + rw add_comm, + exact (-o).tan_oangle_add_right_mul_norm_of_oangle_eq_pi_div_two h +end + +/-- A side of a right-angled triangle divided by the cosine of the adjacent angle equals the +hypotenuse. -/ +lemma norm_div_cos_oangle_add_right_of_oangle_eq_pi_div_two {x y : V} + (h : o.oangle x y = ↑(π / 2)) : ‖x‖ / real.angle.cos (o.oangle x (x + y)) = ‖x + y‖ := +begin + have hs : (o.oangle x (x + y)).sign = 1, + { rw [oangle_sign_add_right, h, real.angle.sign_coe_pi_div_two] }, + rw [o.oangle_eq_angle_of_sign_eq_one hs, real.angle.cos_coe, + inner_product_geometry.norm_div_cos_angle_add_of_inner_eq_zero + (o.inner_eq_zero_of_oangle_eq_pi_div_two h) + (or.inl (o.left_ne_zero_of_oangle_eq_pi_div_two h))] +end + +/-- A side of a right-angled triangle divided by the cosine of the adjacent angle equals the +hypotenuse. -/ +lemma norm_div_cos_oangle_add_left_of_oangle_eq_pi_div_two {x y : V} + (h : o.oangle x y = ↑(π / 2)) : ‖y‖ / real.angle.cos (o.oangle (x + y) y) = ‖x + y‖ := +begin + rw [←neg_inj, oangle_rev, ←oangle_neg_orientation_eq_neg, neg_inj] at h ⊢, + rw add_comm, + exact (-o).norm_div_cos_oangle_add_right_of_oangle_eq_pi_div_two h +end + +/-- A side of a right-angled triangle divided by the sine of the opposite angle equals the +hypotenuse. -/ +lemma norm_div_sin_oangle_add_right_of_oangle_eq_pi_div_two {x y : V} + (h : o.oangle x y = ↑(π / 2)) : ‖y‖ / real.angle.sin (o.oangle x (x + y)) = ‖x + y‖ := +begin + have hs : (o.oangle x (x + y)).sign = 1, + { rw [oangle_sign_add_right, h, real.angle.sign_coe_pi_div_two] }, + rw [o.oangle_eq_angle_of_sign_eq_one hs, real.angle.sin_coe, + inner_product_geometry.norm_div_sin_angle_add_of_inner_eq_zero + (o.inner_eq_zero_of_oangle_eq_pi_div_two h) + (or.inr (o.right_ne_zero_of_oangle_eq_pi_div_two h))] +end + +/-- A side of a right-angled triangle divided by the sine of the opposite angle equals the +hypotenuse. -/ +lemma norm_div_sin_oangle_add_left_of_oangle_eq_pi_div_two {x y : V} + (h : o.oangle x y = ↑(π / 2)) : ‖x‖ / real.angle.sin (o.oangle (x + y) y) = ‖x + y‖ := +begin + rw [←neg_inj, oangle_rev, ←oangle_neg_orientation_eq_neg, neg_inj] at h ⊢, + rw add_comm, + exact (-o).norm_div_sin_oangle_add_right_of_oangle_eq_pi_div_two h +end + +/-- A side of a right-angled triangle divided by the tangent of the opposite angle equals the +adjacent side. -/ +lemma norm_div_tan_oangle_add_right_of_oangle_eq_pi_div_two {x y : V} + (h : o.oangle x y = ↑(π / 2)) : ‖y‖ / real.angle.tan (o.oangle x (x + y)) = ‖x‖ := +begin + have hs : (o.oangle x (x + y)).sign = 1, + { rw [oangle_sign_add_right, h, real.angle.sign_coe_pi_div_two] }, + rw [o.oangle_eq_angle_of_sign_eq_one hs, real.angle.tan_coe, + inner_product_geometry.norm_div_tan_angle_add_of_inner_eq_zero + (o.inner_eq_zero_of_oangle_eq_pi_div_two h) + (or.inr (o.right_ne_zero_of_oangle_eq_pi_div_two h))] +end + +/-- A side of a right-angled triangle divided by the tangent of the opposite angle equals the +adjacent side. -/ +lemma norm_div_tan_oangle_add_left_of_oangle_eq_pi_div_two {x y : V} + (h : o.oangle x y = ↑(π / 2)) : ‖x‖ / real.angle.tan (o.oangle (x + y) y) = ‖y‖ := +begin + rw [←neg_inj, oangle_rev, ←oangle_neg_orientation_eq_neg, neg_inj] at h ⊢, + rw add_comm, + exact (-o).norm_div_tan_oangle_add_right_of_oangle_eq_pi_div_two h +end + +/-- An angle in a right-angled triangle expressed using `arccos`, version subtracting vectors. -/ +lemma oangle_sub_right_eq_arccos_of_oangle_eq_pi_div_two {x y : V} (h : o.oangle x y = ↑(π / 2)) : + o.oangle y (y - x) = real.arccos (‖y‖ / ‖y - x‖) := +begin + have hs : (o.oangle y (y - x)).sign = 1, + { rw [oangle_sign_sub_right_swap, h, real.angle.sign_coe_pi_div_two] }, + rw [o.oangle_eq_angle_of_sign_eq_one hs, + inner_product_geometry.angle_sub_eq_arccos_of_inner_eq_zero + (o.inner_rev_eq_zero_of_oangle_eq_pi_div_two h)] +end + +/-- An angle in a right-angled triangle expressed using `arccos`, version subtracting vectors. -/ +lemma oangle_sub_left_eq_arccos_of_oangle_eq_pi_div_two {x y : V} (h : o.oangle x y = ↑(π / 2)) : + o.oangle (x - y) x = real.arccos (‖x‖ / ‖x - y‖) := +begin + rw [←neg_inj, oangle_rev, ←oangle_neg_orientation_eq_neg, neg_inj] at h ⊢, + exact (-o).oangle_sub_right_eq_arccos_of_oangle_eq_pi_div_two h +end + +/-- An angle in a right-angled triangle expressed using `arcsin`, version subtracting vectors. -/ +lemma oangle_sub_right_eq_arcsin_of_oangle_eq_pi_div_two {x y : V} (h : o.oangle x y = ↑(π / 2)) : + o.oangle y (y - x) = real.arcsin (‖x‖ / ‖y - x‖) := +begin + have hs : (o.oangle y (y - x)).sign = 1, + { rw [oangle_sign_sub_right_swap, h, real.angle.sign_coe_pi_div_two] }, + rw [o.oangle_eq_angle_of_sign_eq_one hs, + inner_product_geometry.angle_sub_eq_arcsin_of_inner_eq_zero + (o.inner_rev_eq_zero_of_oangle_eq_pi_div_two h) + (or.inl (o.right_ne_zero_of_oangle_eq_pi_div_two h))] +end + +/-- An angle in a right-angled triangle expressed using `arcsin`, version subtracting vectors. -/ +lemma oangle_sub_left_eq_arcsin_of_oangle_eq_pi_div_two {x y : V} (h : o.oangle x y = ↑(π / 2)) : + o.oangle (x - y) x = real.arcsin (‖y‖ / ‖x - y‖) := +begin + rw [←neg_inj, oangle_rev, ←oangle_neg_orientation_eq_neg, neg_inj] at h ⊢, + exact (-o).oangle_sub_right_eq_arcsin_of_oangle_eq_pi_div_two h +end + +/-- An angle in a right-angled triangle expressed using `arctan`, version subtracting vectors. -/ +lemma oangle_sub_right_eq_arctan_of_oangle_eq_pi_div_two {x y : V} (h : o.oangle x y = ↑(π / 2)) : + o.oangle y (y - x) = real.arctan (‖x‖ / ‖y‖) := +begin + have hs : (o.oangle y (y - x)).sign = 1, + { rw [oangle_sign_sub_right_swap, h, real.angle.sign_coe_pi_div_two] }, + rw [o.oangle_eq_angle_of_sign_eq_one hs, + inner_product_geometry.angle_sub_eq_arctan_of_inner_eq_zero + (o.inner_rev_eq_zero_of_oangle_eq_pi_div_two h) + (o.right_ne_zero_of_oangle_eq_pi_div_two h)] +end + +/-- An angle in a right-angled triangle expressed using `arctan`, version subtracting vectors. -/ +lemma oangle_sub_left_eq_arctan_of_oangle_eq_pi_div_two {x y : V} (h : o.oangle x y = ↑(π / 2)) : + o.oangle (x - y) x = real.arctan (‖y‖ / ‖x‖) := +begin + rw [←neg_inj, oangle_rev, ←oangle_neg_orientation_eq_neg, neg_inj] at h ⊢, + exact (-o).oangle_sub_right_eq_arctan_of_oangle_eq_pi_div_two h +end + +/-- The cosine of an angle in a right-angled triangle as a ratio of sides, version subtracting +vectors. -/ +lemma cos_oangle_sub_right_of_oangle_eq_pi_div_two {x y : V} (h : o.oangle x y = ↑(π / 2)) : + real.angle.cos (o.oangle y (y - x)) = ‖y‖ / ‖y - x‖ := +begin + have hs : (o.oangle y (y - x)).sign = 1, + { rw [oangle_sign_sub_right_swap, h, real.angle.sign_coe_pi_div_two] }, + rw [o.oangle_eq_angle_of_sign_eq_one hs, real.angle.cos_coe, + inner_product_geometry.cos_angle_sub_of_inner_eq_zero + (o.inner_rev_eq_zero_of_oangle_eq_pi_div_two h)] +end + +/-- The cosine of an angle in a right-angled triangle as a ratio of sides, version subtracting +vectors. -/ +lemma cos_oangle_sub_left_of_oangle_eq_pi_div_two {x y : V} (h : o.oangle x y = ↑(π / 2)) : + real.angle.cos (o.oangle (x - y) x) = ‖x‖ / ‖x - y‖ := +begin + rw [←neg_inj, oangle_rev, ←oangle_neg_orientation_eq_neg, neg_inj] at h ⊢, + exact (-o).cos_oangle_sub_right_of_oangle_eq_pi_div_two h +end + +/-- The sine of an angle in a right-angled triangle as a ratio of sides, version subtracting +vectors. -/ +lemma sin_oangle_sub_right_of_oangle_eq_pi_div_two {x y : V} (h : o.oangle x y = ↑(π / 2)) : + real.angle.sin (o.oangle y (y - x)) = ‖x‖ / ‖y - x‖ := +begin + have hs : (o.oangle y (y - x)).sign = 1, + { rw [oangle_sign_sub_right_swap, h, real.angle.sign_coe_pi_div_two] }, + rw [o.oangle_eq_angle_of_sign_eq_one hs, real.angle.sin_coe, + inner_product_geometry.sin_angle_sub_of_inner_eq_zero + (o.inner_rev_eq_zero_of_oangle_eq_pi_div_two h) + (or.inl (o.right_ne_zero_of_oangle_eq_pi_div_two h))] +end + +/-- The sine of an angle in a right-angled triangle as a ratio of sides, version subtracting +vectors. -/ +lemma sin_oangle_sub_left_of_oangle_eq_pi_div_two {x y : V} (h : o.oangle x y = ↑(π / 2)) : + real.angle.sin (o.oangle (x - y) x) = ‖y‖ / ‖x - y‖ := +begin + rw [←neg_inj, oangle_rev, ←oangle_neg_orientation_eq_neg, neg_inj] at h ⊢, + exact (-o).sin_oangle_sub_right_of_oangle_eq_pi_div_two h +end + +/-- The tangent of an angle in a right-angled triangle as a ratio of sides, version subtracting +vectors. -/ +lemma tan_oangle_sub_right_of_oangle_eq_pi_div_two {x y : V} (h : o.oangle x y = ↑(π / 2)) : + real.angle.tan (o.oangle y (y - x)) = ‖x‖ / ‖y‖ := +begin + have hs : (o.oangle y (y - x)).sign = 1, + { rw [oangle_sign_sub_right_swap, h, real.angle.sign_coe_pi_div_two] }, + rw [o.oangle_eq_angle_of_sign_eq_one hs, real.angle.tan_coe, + inner_product_geometry.tan_angle_sub_of_inner_eq_zero + (o.inner_rev_eq_zero_of_oangle_eq_pi_div_two h)] +end + +/-- The tangent of an angle in a right-angled triangle as a ratio of sides, version subtracting +vectors. -/ +lemma tan_oangle_sub_left_of_oangle_eq_pi_div_two {x y : V} (h : o.oangle x y = ↑(π / 2)) : + real.angle.tan (o.oangle (x - y) x) = ‖y‖ / ‖x‖ := +begin + rw [←neg_inj, oangle_rev, ←oangle_neg_orientation_eq_neg, neg_inj] at h ⊢, + exact (-o).tan_oangle_sub_right_of_oangle_eq_pi_div_two h +end + +/-- The cosine of an angle in a right-angled triangle multiplied by the hypotenuse equals the +adjacent side, version subtracting vectors. -/ +lemma cos_oangle_sub_right_mul_norm_of_oangle_eq_pi_div_two {x y : V} + (h : o.oangle x y = ↑(π / 2)) : real.angle.cos (o.oangle y (y - x)) * ‖y - x‖ = ‖y‖ := +begin + have hs : (o.oangle y (y - x)).sign = 1, + { rw [oangle_sign_sub_right_swap, h, real.angle.sign_coe_pi_div_two] }, + rw [o.oangle_eq_angle_of_sign_eq_one hs, real.angle.cos_coe, + inner_product_geometry.cos_angle_sub_mul_norm_of_inner_eq_zero + (o.inner_rev_eq_zero_of_oangle_eq_pi_div_two h)] +end + +/-- The cosine of an angle in a right-angled triangle multiplied by the hypotenuse equals the +adjacent side, version subtracting vectors. -/ +lemma cos_oangle_sub_left_mul_norm_of_oangle_eq_pi_div_two {x y : V} + (h : o.oangle x y = ↑(π / 2)) : real.angle.cos (o.oangle (x - y) x) * ‖x - y‖ = ‖x‖ := +begin + rw [←neg_inj, oangle_rev, ←oangle_neg_orientation_eq_neg, neg_inj] at h ⊢, + exact (-o).cos_oangle_sub_right_mul_norm_of_oangle_eq_pi_div_two h +end + +/-- The sine of an angle in a right-angled triangle multiplied by the hypotenuse equals the +opposite side, version subtracting vectors. -/ +lemma sin_oangle_sub_right_mul_norm_of_oangle_eq_pi_div_two {x y : V} + (h : o.oangle x y = ↑(π / 2)) : real.angle.sin (o.oangle y (y - x)) * ‖y - x‖ = ‖x‖ := +begin + have hs : (o.oangle y (y - x)).sign = 1, + { rw [oangle_sign_sub_right_swap, h, real.angle.sign_coe_pi_div_two] }, + rw [o.oangle_eq_angle_of_sign_eq_one hs, real.angle.sin_coe, + inner_product_geometry.sin_angle_sub_mul_norm_of_inner_eq_zero + (o.inner_rev_eq_zero_of_oangle_eq_pi_div_two h)] +end + +/-- The sine of an angle in a right-angled triangle multiplied by the hypotenuse equals the +opposite side, version subtracting vectors. -/ +lemma sin_oangle_sub_left_mul_norm_of_oangle_eq_pi_div_two {x y : V} + (h : o.oangle x y = ↑(π / 2)) : real.angle.sin (o.oangle (x - y) x) * ‖x - y‖ = ‖y‖ := +begin + rw [←neg_inj, oangle_rev, ←oangle_neg_orientation_eq_neg, neg_inj] at h ⊢, + exact (-o).sin_oangle_sub_right_mul_norm_of_oangle_eq_pi_div_two h +end + +/-- The tangent of an angle in a right-angled triangle multiplied by the adjacent side equals +the opposite side, version subtracting vectors. -/ +lemma tan_oangle_sub_right_mul_norm_of_oangle_eq_pi_div_two {x y : V} + (h : o.oangle x y = ↑(π / 2)) : real.angle.tan (o.oangle y (y - x)) * ‖y‖ = ‖x‖ := +begin + have hs : (o.oangle y (y - x)).sign = 1, + { rw [oangle_sign_sub_right_swap, h, real.angle.sign_coe_pi_div_two] }, + rw [o.oangle_eq_angle_of_sign_eq_one hs, real.angle.tan_coe, + inner_product_geometry.tan_angle_sub_mul_norm_of_inner_eq_zero + (o.inner_rev_eq_zero_of_oangle_eq_pi_div_two h) + (or.inl (o.right_ne_zero_of_oangle_eq_pi_div_two h))] +end + +/-- The tangent of an angle in a right-angled triangle multiplied by the adjacent side equals +the opposite side, version subtracting vectors. -/ +lemma tan_oangle_sub_left_mul_norm_of_oangle_eq_pi_div_two {x y : V} + (h : o.oangle x y = ↑(π / 2)) : real.angle.tan (o.oangle (x - y) x) * ‖x‖ = ‖y‖ := +begin + rw [←neg_inj, oangle_rev, ←oangle_neg_orientation_eq_neg, neg_inj] at h ⊢, + exact (-o).tan_oangle_sub_right_mul_norm_of_oangle_eq_pi_div_two h +end + +/-- A side of a right-angled triangle divided by the cosine of the adjacent angle equals the +hypotenuse, version subtracting vectors. -/ +lemma norm_div_cos_oangle_sub_right_of_oangle_eq_pi_div_two {x y : V} + (h : o.oangle x y = ↑(π / 2)) : ‖y‖ / real.angle.cos (o.oangle y (y - x)) = ‖y - x‖ := +begin + have hs : (o.oangle y (y - x)).sign = 1, + { rw [oangle_sign_sub_right_swap, h, real.angle.sign_coe_pi_div_two] }, + rw [o.oangle_eq_angle_of_sign_eq_one hs, real.angle.cos_coe, + inner_product_geometry.norm_div_cos_angle_sub_of_inner_eq_zero + (o.inner_rev_eq_zero_of_oangle_eq_pi_div_two h) + (or.inl (o.right_ne_zero_of_oangle_eq_pi_div_two h))] +end + +/-- A side of a right-angled triangle divided by the cosine of the adjacent angle equals the +hypotenuse, version subtracting vectors. -/ +lemma norm_div_cos_oangle_sub_left_of_oangle_eq_pi_div_two {x y : V} + (h : o.oangle x y = ↑(π / 2)) : ‖x‖ / real.angle.cos (o.oangle (x - y) x) = ‖x - y‖ := +begin + rw [←neg_inj, oangle_rev, ←oangle_neg_orientation_eq_neg, neg_inj] at h ⊢, + exact (-o).norm_div_cos_oangle_sub_right_of_oangle_eq_pi_div_two h +end + +/-- A side of a right-angled triangle divided by the sine of the opposite angle equals the +hypotenuse, version subtracting vectors. -/ +lemma norm_div_sin_oangle_sub_right_of_oangle_eq_pi_div_two {x y : V} + (h : o.oangle x y = ↑(π / 2)) : ‖x‖ / real.angle.sin (o.oangle y (y - x)) = ‖y - x‖ := +begin + have hs : (o.oangle y (y - x)).sign = 1, + { rw [oangle_sign_sub_right_swap, h, real.angle.sign_coe_pi_div_two] }, + rw [o.oangle_eq_angle_of_sign_eq_one hs, real.angle.sin_coe, + inner_product_geometry.norm_div_sin_angle_sub_of_inner_eq_zero + (o.inner_rev_eq_zero_of_oangle_eq_pi_div_two h) + (or.inr (o.left_ne_zero_of_oangle_eq_pi_div_two h))] +end + +/-- A side of a right-angled triangle divided by the sine of the opposite angle equals the +hypotenuse, version subtracting vectors. -/ +lemma norm_div_sin_oangle_sub_left_of_oangle_eq_pi_div_two {x y : V} + (h : o.oangle x y = ↑(π / 2)) : ‖y‖ / real.angle.sin (o.oangle (x - y) x) = ‖x - y‖ := +begin + rw [←neg_inj, oangle_rev, ←oangle_neg_orientation_eq_neg, neg_inj] at h ⊢, + exact (-o).norm_div_sin_oangle_sub_right_of_oangle_eq_pi_div_two h +end + +/-- A side of a right-angled triangle divided by the tangent of the opposite angle equals the +adjacent side, version subtracting vectors. -/ +lemma norm_div_tan_oangle_sub_right_of_oangle_eq_pi_div_two {x y : V} + (h : o.oangle x y = ↑(π / 2)) : ‖x‖ / real.angle.tan (o.oangle y (y - x)) = ‖y‖ := +begin + have hs : (o.oangle y (y - x)).sign = 1, + { rw [oangle_sign_sub_right_swap, h, real.angle.sign_coe_pi_div_two] }, + rw [o.oangle_eq_angle_of_sign_eq_one hs, real.angle.tan_coe, + inner_product_geometry.norm_div_tan_angle_sub_of_inner_eq_zero + (o.inner_rev_eq_zero_of_oangle_eq_pi_div_two h) + (or.inr (o.left_ne_zero_of_oangle_eq_pi_div_two h))] +end + +/-- A side of a right-angled triangle divided by the tangent of the opposite angle equals the +adjacent side, version subtracting vectors. -/ +lemma norm_div_tan_oangle_sub_left_of_oangle_eq_pi_div_two {x y : V} + (h : o.oangle x y = ↑(π / 2)) : ‖y‖ / real.angle.tan (o.oangle (x - y) x) = ‖x‖ := +begin + rw [←neg_inj, oangle_rev, ←oangle_neg_orientation_eq_neg, neg_inj] at h ⊢, + exact (-o).norm_div_tan_oangle_sub_right_of_oangle_eq_pi_div_two h +end + +/-- An angle in a right-angled triangle expressed using `arctan`, where one side is a multiple +of a rotation of another by `π / 2`. -/ +lemma oangle_add_right_smul_rotation_pi_div_two {x : V} (h : x ≠ 0) (r : ℝ) : + o.oangle x (x + r • o.rotation (π / 2 : ℝ) x) = real.arctan r := +begin + rcases lt_trichotomy r 0 with hr | rfl | hr, + { have ha : o.oangle x (r • o.rotation (π / 2 : ℝ) x) = -(π / 2 : ℝ), + { rw [o.oangle_smul_right_of_neg _ _ hr, o.oangle_neg_right h, + o.oangle_rotation_self_right h, ←sub_eq_zero, add_comm, sub_neg_eq_add, + ←real.angle.coe_add, ←real.angle.coe_add, add_assoc, add_halves, ←two_mul, + real.angle.coe_two_pi], + simpa using h }, + rw [←neg_inj, ←oangle_neg_orientation_eq_neg, neg_neg] at ha, + rw [←neg_inj, oangle_rev, ←oangle_neg_orientation_eq_neg, neg_inj, oangle_rev, + (-o).oangle_add_right_eq_arctan_of_oangle_eq_pi_div_two ha, norm_smul, + linear_isometry_equiv.norm_map, mul_div_assoc, div_self (norm_ne_zero_iff.2 h), mul_one, + real.norm_eq_abs, abs_of_neg hr, real.arctan_neg, real.angle.coe_neg, neg_neg] }, + { rw [zero_smul, add_zero, oangle_self, real.arctan_zero, real.angle.coe_zero] }, + { have ha : o.oangle x (r • o.rotation (π / 2 : ℝ) x) = (π / 2 : ℝ), + { rw [o.oangle_smul_right_of_pos _ _ hr, o.oangle_rotation_self_right h] }, + rw [o.oangle_add_right_eq_arctan_of_oangle_eq_pi_div_two ha, norm_smul, + linear_isometry_equiv.norm_map, mul_div_assoc, div_self (norm_ne_zero_iff.2 h), mul_one, + real.norm_eq_abs, abs_of_pos hr] } +end + +/-- An angle in a right-angled triangle expressed using `arctan`, where one side is a multiple +of a rotation of another by `π / 2`. -/ +lemma oangle_add_left_smul_rotation_pi_div_two {x : V} (h : x ≠ 0) (r : ℝ) : + o.oangle (x + r • o.rotation (π / 2 : ℝ) x) (r • o.rotation (π / 2 : ℝ) x) = real.arctan r⁻¹ := +begin + by_cases hr : r = 0, { simp [hr] }, + rw [←neg_inj, oangle_rev, ←oangle_neg_orientation_eq_neg, neg_inj, + ←neg_neg ((π / 2 : ℝ) : real.angle), ←rotation_neg_orientation_eq_neg, add_comm], + have hx : x = r⁻¹ • ((-o).rotation (π / 2 : ℝ) (r • ((-o).rotation (-(π / 2 : ℝ)) x))), + { simp [hr] }, + nth_rewrite 2 hx, + refine (-o).oangle_add_right_smul_rotation_pi_div_two _ _, + simp [hr, h] +end + +/-- The tangent of an angle in a right-angled triangle, where one side is a multiple of a +rotation of another by `π / 2`. -/ +lemma tan_oangle_add_right_smul_rotation_pi_div_two {x : V} (h : x ≠ 0) (r : ℝ) : + real.angle.tan (o.oangle x (x + r • o.rotation (π / 2 : ℝ) x)) = r := +by rw [o.oangle_add_right_smul_rotation_pi_div_two h, real.angle.tan_coe, real.tan_arctan] + +/-- The tangent of an angle in a right-angled triangle, where one side is a multiple of a +rotation of another by `π / 2`. -/ +lemma tan_oangle_add_left_smul_rotation_pi_div_two {x : V} (h : x ≠ 0) (r : ℝ) : + real.angle.tan (o.oangle (x + r • o.rotation (π / 2 : ℝ) x) (r • o.rotation (π / 2 : ℝ) x)) = + r⁻¹ := +by rw [o.oangle_add_left_smul_rotation_pi_div_two h, real.angle.tan_coe, real.tan_arctan] + +/-- An angle in a right-angled triangle expressed using `arctan`, where one side is a multiple +of a rotation of another by `π / 2`, version subtracting vectors. -/ +lemma oangle_sub_right_smul_rotation_pi_div_two {x : V} (h : x ≠ 0) (r : ℝ) : + o.oangle (r • o.rotation (π / 2 : ℝ) x) (r • o.rotation (π / 2 : ℝ) x - x) = real.arctan r⁻¹ := +begin + by_cases hr : r = 0, { simp [hr] }, + have hx : -x = r⁻¹ • (o.rotation (π / 2 : ℝ) (r • (o.rotation (π / 2 : ℝ) x))), + { simp [hr, ←real.angle.coe_add] }, + rw [sub_eq_add_neg, hx, o.oangle_add_right_smul_rotation_pi_div_two], + simpa [hr] using h +end + +/-- An angle in a right-angled triangle expressed using `arctan`, where one side is a multiple +of a rotation of another by `π / 2`, version subtracting vectors. -/ +lemma oangle_sub_left_smul_rotation_pi_div_two {x : V} (h : x ≠ 0) (r : ℝ) : + o.oangle (x - r • o.rotation (π / 2 : ℝ) x) x = real.arctan r := +begin + by_cases hr : r = 0, { simp [hr] }, + have hx : x = r⁻¹ • (o.rotation (π / 2 : ℝ) (-(r • (o.rotation (π / 2 : ℝ) x)))), + { simp [hr, ←real.angle.coe_add] }, + rw [sub_eq_add_neg, add_comm], + nth_rewrite 2 hx, + nth_rewrite 1 hx, + rw [o.oangle_add_left_smul_rotation_pi_div_two, inv_inv], + simpa [hr] using h +end + +end orientation + +namespace euclidean_geometry + +open finite_dimensional + +variables {V : Type*} {P : Type*} [inner_product_space ℝ V] [metric_space P] +variables [normed_add_torsor V P] [hd2 : fact (finrank ℝ V = 2)] [module.oriented ℝ V (fin 2)] +include hd2 + +/-- An angle in a right-angled triangle expressed using `arccos`. -/ +lemma oangle_right_eq_arccos_of_oangle_eq_pi_div_two {p₁ p₂ p₃ : P} (h : ∡ p₁ p₂ p₃ = ↑(π / 2)) : + ∡ p₂ p₃ p₁ = real.arccos (dist p₃ p₂ / dist p₁ p₃) := +begin + have hs : (∡ p₂ p₃ p₁).sign = 1, { rw [oangle_rotate_sign, h, real.angle.sign_coe_pi_div_two] }, + rw [oangle_eq_angle_of_sign_eq_one hs, + angle_eq_arccos_of_angle_eq_pi_div_two (angle_eq_pi_div_two_of_oangle_eq_pi_div_two h)] +end + +/-- An angle in a right-angled triangle expressed using `arccos`. -/ +lemma oangle_left_eq_arccos_of_oangle_eq_pi_div_two {p₁ p₂ p₃ : P} (h : ∡ p₁ p₂ p₃ = ↑(π / 2)) : + ∡ p₃ p₁ p₂ = real.arccos (dist p₁ p₂ / dist p₁ p₃) := +begin + have hs : (∡ p₃ p₁ p₂).sign = 1, { rw [←oangle_rotate_sign, h, real.angle.sign_coe_pi_div_two] }, + rw [oangle_eq_angle_of_sign_eq_one hs, angle_comm, + angle_eq_arccos_of_angle_eq_pi_div_two (angle_rev_eq_pi_div_two_of_oangle_eq_pi_div_two h), + dist_comm p₁ p₃] +end + +/-- An angle in a right-angled triangle expressed using `arcsin`. -/ +lemma oangle_right_eq_arcsin_of_oangle_eq_pi_div_two {p₁ p₂ p₃ : P} (h : ∡ p₁ p₂ p₃ = ↑(π / 2)) : + ∡ p₂ p₃ p₁ = real.arcsin (dist p₁ p₂ / dist p₁ p₃) := +begin + have hs : (∡ p₂ p₃ p₁).sign = 1, { rw [oangle_rotate_sign, h, real.angle.sign_coe_pi_div_two] }, + rw [oangle_eq_angle_of_sign_eq_one hs, + angle_eq_arcsin_of_angle_eq_pi_div_two (angle_eq_pi_div_two_of_oangle_eq_pi_div_two h) + (or.inl (left_ne_of_oangle_eq_pi_div_two h))] +end + +/-- An angle in a right-angled triangle expressed using `arcsin`. -/ +lemma oangle_left_eq_arcsin_of_oangle_eq_pi_div_two {p₁ p₂ p₃ : P} (h : ∡ p₁ p₂ p₃ = ↑(π / 2)) : + ∡ p₃ p₁ p₂ = real.arcsin (dist p₃ p₂ / dist p₁ p₃) := +begin + have hs : (∡ p₃ p₁ p₂).sign = 1, { rw [←oangle_rotate_sign, h, real.angle.sign_coe_pi_div_two] }, + rw [oangle_eq_angle_of_sign_eq_one hs, angle_comm, + angle_eq_arcsin_of_angle_eq_pi_div_two (angle_rev_eq_pi_div_two_of_oangle_eq_pi_div_two h) + (or.inr (left_ne_of_oangle_eq_pi_div_two h)), + dist_comm p₁ p₃] +end + +/-- An angle in a right-angled triangle expressed using `arctan`. -/ +lemma oangle_right_eq_arctan_of_oangle_eq_pi_div_two {p₁ p₂ p₃ : P} (h : ∡ p₁ p₂ p₃ = ↑(π / 2)) : + ∡ p₂ p₃ p₁ = real.arctan (dist p₁ p₂ / dist p₃ p₂) := +begin + have hs : (∡ p₂ p₃ p₁).sign = 1, { rw [oangle_rotate_sign, h, real.angle.sign_coe_pi_div_two] }, + rw [oangle_eq_angle_of_sign_eq_one hs, + angle_eq_arctan_of_angle_eq_pi_div_two (angle_eq_pi_div_two_of_oangle_eq_pi_div_two h) + (right_ne_of_oangle_eq_pi_div_two h)] +end + +/-- An angle in a right-angled triangle expressed using `arctan`. -/ +lemma oangle_left_eq_arctan_of_oangle_eq_pi_div_two {p₁ p₂ p₃ : P} (h : ∡ p₁ p₂ p₃ = ↑(π / 2)) : + ∡ p₃ p₁ p₂ = real.arctan (dist p₃ p₂ / dist p₁ p₂) := +begin + have hs : (∡ p₃ p₁ p₂).sign = 1, { rw [←oangle_rotate_sign, h, real.angle.sign_coe_pi_div_two] }, + rw [oangle_eq_angle_of_sign_eq_one hs, angle_comm, + angle_eq_arctan_of_angle_eq_pi_div_two (angle_rev_eq_pi_div_two_of_oangle_eq_pi_div_two h) + (left_ne_of_oangle_eq_pi_div_two h)] +end + +/-- The cosine of an angle in a right-angled triangle as a ratio of sides. -/ +lemma cos_oangle_right_of_oangle_eq_pi_div_two {p₁ p₂ p₃ : P} (h : ∡ p₁ p₂ p₃ = ↑(π / 2)) : + real.angle.cos (∡ p₂ p₃ p₁) = dist p₃ p₂ / dist p₁ p₃ := +begin + have hs : (∡ p₂ p₃ p₁).sign = 1, { rw [oangle_rotate_sign, h, real.angle.sign_coe_pi_div_two] }, + rw [oangle_eq_angle_of_sign_eq_one hs, real.angle.cos_coe, + cos_angle_of_angle_eq_pi_div_two (angle_eq_pi_div_two_of_oangle_eq_pi_div_two h)] +end + +/-- The cosine of an angle in a right-angled triangle as a ratio of sides. -/ +lemma cos_oangle_left_of_oangle_eq_pi_div_two {p₁ p₂ p₃ : P} (h : ∡ p₁ p₂ p₃ = ↑(π / 2)) : + real.angle.cos (∡ p₃ p₁ p₂) = dist p₁ p₂ / dist p₁ p₃ := +begin + have hs : (∡ p₃ p₁ p₂).sign = 1, { rw [←oangle_rotate_sign, h, real.angle.sign_coe_pi_div_two] }, + rw [oangle_eq_angle_of_sign_eq_one hs, angle_comm, real.angle.cos_coe, + cos_angle_of_angle_eq_pi_div_two (angle_rev_eq_pi_div_two_of_oangle_eq_pi_div_two h), + dist_comm p₁ p₃] +end + +/-- The sine of an angle in a right-angled triangle as a ratio of sides. -/ +lemma sin_oangle_right_of_oangle_eq_pi_div_two {p₁ p₂ p₃ : P} (h : ∡ p₁ p₂ p₃ = ↑(π / 2)) : + real.angle.sin (∡ p₂ p₃ p₁) = dist p₁ p₂ / dist p₁ p₃ := +begin + have hs : (∡ p₂ p₃ p₁).sign = 1, { rw [oangle_rotate_sign, h, real.angle.sign_coe_pi_div_two] }, + rw [oangle_eq_angle_of_sign_eq_one hs, real.angle.sin_coe, + sin_angle_of_angle_eq_pi_div_two (angle_eq_pi_div_two_of_oangle_eq_pi_div_two h) + (or.inl (left_ne_of_oangle_eq_pi_div_two h))] +end + +/-- The sine of an angle in a right-angled triangle as a ratio of sides. -/ +lemma sin_oangle_left_of_oangle_eq_pi_div_two {p₁ p₂ p₃ : P} (h : ∡ p₁ p₂ p₃ = ↑(π / 2)) : + real.angle.sin (∡ p₃ p₁ p₂) = dist p₃ p₂ / dist p₁ p₃ := +begin + have hs : (∡ p₃ p₁ p₂).sign = 1, { rw [←oangle_rotate_sign, h, real.angle.sign_coe_pi_div_two] }, + rw [oangle_eq_angle_of_sign_eq_one hs, angle_comm, real.angle.sin_coe, + sin_angle_of_angle_eq_pi_div_two (angle_rev_eq_pi_div_two_of_oangle_eq_pi_div_two h) + (or.inr (left_ne_of_oangle_eq_pi_div_two h)), + dist_comm p₁ p₃] +end + +/-- The tangent of an angle in a right-angled triangle as a ratio of sides. -/ +lemma tan_oangle_right_of_oangle_eq_pi_div_two {p₁ p₂ p₃ : P} (h : ∡ p₁ p₂ p₃ = ↑(π / 2)) : + real.angle.tan (∡ p₂ p₃ p₁) = dist p₁ p₂ / dist p₃ p₂ := +begin + have hs : (∡ p₂ p₃ p₁).sign = 1, { rw [oangle_rotate_sign, h, real.angle.sign_coe_pi_div_two] }, + rw [oangle_eq_angle_of_sign_eq_one hs, real.angle.tan_coe, + tan_angle_of_angle_eq_pi_div_two (angle_eq_pi_div_two_of_oangle_eq_pi_div_two h)] +end + +/-- The tangent of an angle in a right-angled triangle as a ratio of sides. -/ +lemma tan_oangle_left_of_oangle_eq_pi_div_two {p₁ p₂ p₃ : P} (h : ∡ p₁ p₂ p₃ = ↑(π / 2)) : + real.angle.tan (∡ p₃ p₁ p₂) = dist p₃ p₂ / dist p₁ p₂ := +begin + have hs : (∡ p₃ p₁ p₂).sign = 1, { rw [←oangle_rotate_sign, h, real.angle.sign_coe_pi_div_two] }, + rw [oangle_eq_angle_of_sign_eq_one hs, angle_comm, real.angle.tan_coe, + tan_angle_of_angle_eq_pi_div_two (angle_rev_eq_pi_div_two_of_oangle_eq_pi_div_two h)] +end + +/-- The cosine of an angle in a right-angled triangle multiplied by the hypotenuse equals the +adjacent side. -/ +lemma cos_oangle_right_mul_dist_of_oangle_eq_pi_div_two {p₁ p₂ p₃ : P} (h : ∡ p₁ p₂ p₃ = ↑(π / 2)) : + real.angle.cos (∡ p₂ p₃ p₁) * dist p₁ p₃ = dist p₃ p₂ := +begin + have hs : (∡ p₂ p₃ p₁).sign = 1, { rw [oangle_rotate_sign, h, real.angle.sign_coe_pi_div_two] }, + rw [oangle_eq_angle_of_sign_eq_one hs, real.angle.cos_coe, + cos_angle_mul_dist_of_angle_eq_pi_div_two (angle_eq_pi_div_two_of_oangle_eq_pi_div_two h)] +end + +/-- The cosine of an angle in a right-angled triangle multiplied by the hypotenuse equals the +adjacent side. -/ +lemma cos_oangle_left_mul_dist_of_oangle_eq_pi_div_two {p₁ p₂ p₃ : P} (h : ∡ p₁ p₂ p₃ = ↑(π / 2)) : + real.angle.cos (∡ p₃ p₁ p₂) * dist p₁ p₃ = dist p₁ p₂ := +begin + have hs : (∡ p₃ p₁ p₂).sign = 1, { rw [←oangle_rotate_sign, h, real.angle.sign_coe_pi_div_two] }, + rw [oangle_eq_angle_of_sign_eq_one hs, angle_comm, real.angle.cos_coe, dist_comm p₁ p₃, + cos_angle_mul_dist_of_angle_eq_pi_div_two (angle_rev_eq_pi_div_two_of_oangle_eq_pi_div_two h)] +end + +/-- The sine of an angle in a right-angled triangle multiplied by the hypotenuse equals the +opposite side. -/ +lemma sin_oangle_right_mul_dist_of_oangle_eq_pi_div_two {p₁ p₂ p₃ : P} (h : ∡ p₁ p₂ p₃ = ↑(π / 2)) : + real.angle.sin (∡ p₂ p₃ p₁) * dist p₁ p₃ = dist p₁ p₂ := +begin + have hs : (∡ p₂ p₃ p₁).sign = 1, { rw [oangle_rotate_sign, h, real.angle.sign_coe_pi_div_two] }, + rw [oangle_eq_angle_of_sign_eq_one hs, real.angle.sin_coe, + sin_angle_mul_dist_of_angle_eq_pi_div_two (angle_eq_pi_div_two_of_oangle_eq_pi_div_two h)] +end + +/-- The sine of an angle in a right-angled triangle multiplied by the hypotenuse equals the +opposite side. -/ +lemma sin_oangle_left_mul_dist_of_oangle_eq_pi_div_two {p₁ p₂ p₃ : P} (h : ∡ p₁ p₂ p₃ = ↑(π / 2)) : + real.angle.sin (∡ p₃ p₁ p₂) * dist p₁ p₃ = dist p₃ p₂ := +begin + have hs : (∡ p₃ p₁ p₂).sign = 1, { rw [←oangle_rotate_sign, h, real.angle.sign_coe_pi_div_two] }, + rw [oangle_eq_angle_of_sign_eq_one hs, angle_comm, real.angle.sin_coe, dist_comm p₁ p₃, + sin_angle_mul_dist_of_angle_eq_pi_div_two (angle_rev_eq_pi_div_two_of_oangle_eq_pi_div_two h)] +end + +/-- The tangent of an angle in a right-angled triangle multiplied by the adjacent side equals +the opposite side. -/ +lemma tan_oangle_right_mul_dist_of_oangle_eq_pi_div_two {p₁ p₂ p₃ : P} (h : ∡ p₁ p₂ p₃ = ↑(π / 2)) : + real.angle.tan (∡ p₂ p₃ p₁) * dist p₃ p₂ = dist p₁ p₂ := +begin + have hs : (∡ p₂ p₃ p₁).sign = 1, { rw [oangle_rotate_sign, h, real.angle.sign_coe_pi_div_two] }, + rw [oangle_eq_angle_of_sign_eq_one hs, real.angle.tan_coe, + tan_angle_mul_dist_of_angle_eq_pi_div_two (angle_eq_pi_div_two_of_oangle_eq_pi_div_two h) + (or.inr (right_ne_of_oangle_eq_pi_div_two h))] +end + +/-- The tangent of an angle in a right-angled triangle multiplied by the adjacent side equals +the opposite side. -/ +lemma tan_oangle_left_mul_dist_of_oangle_eq_pi_div_two {p₁ p₂ p₃ : P} (h : ∡ p₁ p₂ p₃ = ↑(π / 2)) : + real.angle.tan (∡ p₃ p₁ p₂) * dist p₁ p₂ = dist p₃ p₂ := +begin + have hs : (∡ p₃ p₁ p₂).sign = 1, { rw [←oangle_rotate_sign, h, real.angle.sign_coe_pi_div_two] }, + rw [oangle_eq_angle_of_sign_eq_one hs, angle_comm, real.angle.tan_coe, + tan_angle_mul_dist_of_angle_eq_pi_div_two (angle_rev_eq_pi_div_two_of_oangle_eq_pi_div_two h) + (or.inr (left_ne_of_oangle_eq_pi_div_two h))] +end + +/-- A side of a right-angled triangle divided by the cosine of the adjacent angle equals the +hypotenuse. -/ +lemma dist_div_cos_oangle_right_of_oangle_eq_pi_div_two {p₁ p₂ p₃ : P} (h : ∡ p₁ p₂ p₃ = ↑(π / 2)) : + dist p₃ p₂ / real.angle.cos (∡ p₂ p₃ p₁) = dist p₁ p₃ := +begin + have hs : (∡ p₂ p₃ p₁).sign = 1, { rw [oangle_rotate_sign, h, real.angle.sign_coe_pi_div_two] }, + rw [oangle_eq_angle_of_sign_eq_one hs, real.angle.cos_coe, + dist_div_cos_angle_of_angle_eq_pi_div_two (angle_eq_pi_div_two_of_oangle_eq_pi_div_two h) + (or.inr (right_ne_of_oangle_eq_pi_div_two h))] +end + +/-- A side of a right-angled triangle divided by the cosine of the adjacent angle equals the +hypotenuse. -/ +lemma dist_div_cos_oangle_left_of_oangle_eq_pi_div_two {p₁ p₂ p₃ : P} (h : ∡ p₁ p₂ p₃ = ↑(π / 2)) : + dist p₁ p₂ / real.angle.cos (∡ p₃ p₁ p₂) = dist p₁ p₃ := +begin + have hs : (∡ p₃ p₁ p₂).sign = 1, { rw [←oangle_rotate_sign, h, real.angle.sign_coe_pi_div_two] }, + rw [oangle_eq_angle_of_sign_eq_one hs, angle_comm, real.angle.cos_coe, dist_comm p₁ p₃, + dist_div_cos_angle_of_angle_eq_pi_div_two (angle_rev_eq_pi_div_two_of_oangle_eq_pi_div_two h) + (or.inr (left_ne_of_oangle_eq_pi_div_two h))] +end + +/-- A side of a right-angled triangle divided by the sine of the opposite angle equals the +hypotenuse. -/ +lemma dist_div_sin_oangle_right_of_oangle_eq_pi_div_two {p₁ p₂ p₃ : P} (h : ∡ p₁ p₂ p₃ = ↑(π / 2)) : + dist p₁ p₂ / real.angle.sin (∡ p₂ p₃ p₁) = dist p₁ p₃ := +begin + have hs : (∡ p₂ p₃ p₁).sign = 1, { rw [oangle_rotate_sign, h, real.angle.sign_coe_pi_div_two] }, + rw [oangle_eq_angle_of_sign_eq_one hs, real.angle.sin_coe, + dist_div_sin_angle_of_angle_eq_pi_div_two (angle_eq_pi_div_two_of_oangle_eq_pi_div_two h) + (or.inl (left_ne_of_oangle_eq_pi_div_two h))] +end + +/-- A side of a right-angled triangle divided by the sine of the opposite angle equals the +hypotenuse. -/ +lemma dist_div_sin_oangle_left_of_oangle_eq_pi_div_two {p₁ p₂ p₃ : P} (h : ∡ p₁ p₂ p₃ = ↑(π / 2)) : + dist p₃ p₂ / real.angle.sin (∡ p₃ p₁ p₂) = dist p₁ p₃ := +begin + have hs : (∡ p₃ p₁ p₂).sign = 1, { rw [←oangle_rotate_sign, h, real.angle.sign_coe_pi_div_two] }, + rw [oangle_eq_angle_of_sign_eq_one hs, angle_comm, real.angle.sin_coe, dist_comm p₁ p₃, + dist_div_sin_angle_of_angle_eq_pi_div_two (angle_rev_eq_pi_div_two_of_oangle_eq_pi_div_two h) + (or.inl (right_ne_of_oangle_eq_pi_div_two h))] +end + +/-- A side of a right-angled triangle divided by the tangent of the opposite angle equals the +adjacent side. -/ +lemma dist_div_tan_oangle_right_of_oangle_eq_pi_div_two {p₁ p₂ p₃ : P} (h : ∡ p₁ p₂ p₃ = ↑(π / 2)) : + dist p₁ p₂ / real.angle.tan (∡ p₂ p₃ p₁) = dist p₃ p₂ := +begin + have hs : (∡ p₂ p₃ p₁).sign = 1, { rw [oangle_rotate_sign, h, real.angle.sign_coe_pi_div_two] }, + rw [oangle_eq_angle_of_sign_eq_one hs, real.angle.tan_coe, + dist_div_tan_angle_of_angle_eq_pi_div_two (angle_eq_pi_div_two_of_oangle_eq_pi_div_two h) + (or.inl (left_ne_of_oangle_eq_pi_div_two h))] +end + +/-- A side of a right-angled triangle divided by the tangent of the opposite angle equals the +adjacent side. -/ +lemma dist_div_tan_oangle_left_of_oangle_eq_pi_div_two {p₁ p₂ p₃ : P} (h : ∡ p₁ p₂ p₃ = ↑(π / 2)) : + dist p₃ p₂ / real.angle.tan (∡ p₃ p₁ p₂) = dist p₁ p₂ := +begin + have hs : (∡ p₃ p₁ p₂).sign = 1, { rw [←oangle_rotate_sign, h, real.angle.sign_coe_pi_div_two] }, + rw [oangle_eq_angle_of_sign_eq_one hs, angle_comm, real.angle.tan_coe, + dist_div_tan_angle_of_angle_eq_pi_div_two (angle_rev_eq_pi_div_two_of_oangle_eq_pi_div_two h) + (or.inl (right_ne_of_oangle_eq_pi_div_two h))] +end + +end euclidean_geometry From ea9c24f633dbd90e37891c05bac8f5027986853d Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Wed, 21 Dec 2022 06:49:18 +0000 Subject: [PATCH 0084/1291] feat(geometry/euclidean/basic): second intersection of a line and a sphere (#17843) A very common geometrical construction is taking the second intersection of a line and a sphere (where one intersection point is given); we already have a lemma `dist_smul_vadd_eq_dist` that gives a characterization of that point. Add an explicit definition `sphere.second_inter` for this construction and associated API. --- src/geometry/euclidean/basic.lean | 214 ++++++++++++++++++++++++++++++ 1 file changed, 214 insertions(+) diff --git a/src/geometry/euclidean/basic.lean b/src/geometry/euclidean/basic.lean index 117be50a146e7..e3e284c12be3b 100644 --- a/src/geometry/euclidean/basic.lean +++ b/src/geometry/euclidean/basic.lean @@ -3,6 +3,7 @@ Copyright (c) 2020 Joseph Myers. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Joseph Myers, Manuel Candales -/ +import analysis.convex.strict_convex_between import analysis.inner_product_space.projection import algebra.quadratic_discriminant @@ -892,6 +893,219 @@ lemma eq_of_mem_sphere_of_mem_sphere_of_finrank_eq_two [finite_dimensional ℝ V eq_of_dist_eq_of_dist_eq_of_finrank_eq_two hd ((sphere.center_ne_iff_ne_of_mem hps₁ hps₂).2 hs) hp hp₁s₁ hp₂s₁ hps₁ hp₁s₂ hp₂s₂ hps₂ +/-- Given a point on a sphere and a point not outside it, the inner product between the +difference of those points and the radius vector is positive unless the points are equal. -/ +lemma inner_pos_or_eq_of_dist_le_radius {s : sphere P} {p₁ p₂ : P} (hp₁ : p₁ ∈ s) + (hp₂ : dist p₂ s.center ≤ s.radius) : 0 < ⟪p₁ -ᵥ p₂, p₁ -ᵥ s.center⟫ ∨ p₁ = p₂ := +begin + by_cases h : p₁ = p₂, { exact or.inr h }, + refine or.inl _, + rw mem_sphere at hp₁, + rw [←vsub_sub_vsub_cancel_right p₁ p₂ s.center, inner_sub_left, + real_inner_self_eq_norm_mul_norm/-, ←dist_eq_norm_vsub, hp₁-/, sub_pos], + refine lt_of_le_of_ne + ((real_inner_le_norm _ _).trans (mul_le_mul_of_nonneg_right _ (norm_nonneg _))) + _, + { rwa [←dist_eq_norm_vsub, ←dist_eq_norm_vsub, hp₁] }, + { rcases hp₂.lt_or_eq with hp₂' | hp₂', + { refine ((real_inner_le_norm _ _).trans_lt (mul_lt_mul_of_pos_right _ _)).ne, + { rwa [←hp₁, @dist_eq_norm_vsub V, @dist_eq_norm_vsub V] at hp₂' }, + { rw [norm_pos_iff, vsub_ne_zero], + rintro rfl, + rw ←hp₁ at hp₂', + refine (dist_nonneg.not_lt : ¬dist p₂ s.center < 0) _, + simpa using hp₂' } }, + { rw [←hp₁, @dist_eq_norm_vsub V, @dist_eq_norm_vsub V] at hp₂', + nth_rewrite 0 ←hp₂', + rw [ne.def, inner_eq_norm_mul_iff_real, hp₂', ←sub_eq_zero, ←smul_sub, + vsub_sub_vsub_cancel_right, ←ne.def, smul_ne_zero_iff, vsub_ne_zero, + and_iff_left (ne.symm h), norm_ne_zero_iff, vsub_ne_zero], + rintro rfl, + refine h (eq.symm _), + simpa using hp₂' } } +end + +/-- Given a point on a sphere and a point not outside it, the inner product between the +difference of those points and the radius vector is nonnegative. -/ +lemma inner_nonneg_of_dist_le_radius {s : sphere P} {p₁ p₂ : P} (hp₁ : p₁ ∈ s) + (hp₂ : dist p₂ s.center ≤ s.radius) : 0 ≤ ⟪p₁ -ᵥ p₂, p₁ -ᵥ s.center⟫ := +begin + rcases inner_pos_or_eq_of_dist_le_radius hp₁ hp₂ with h | rfl, + { exact h.le }, + { simp } +end + +/-- Given a point on a sphere and a point inside it, the inner product between the difference of +those points and the radius vector is positive. -/ +lemma inner_pos_of_dist_lt_radius {s : sphere P} {p₁ p₂ : P} (hp₁ : p₁ ∈ s) + (hp₂ : dist p₂ s.center < s.radius) : 0 < ⟪p₁ -ᵥ p₂, p₁ -ᵥ s.center⟫ := +begin + by_cases h : p₁ = p₂, + { rw [h, mem_sphere] at hp₁, + exact false.elim (hp₂.ne hp₁) }, + exact (inner_pos_or_eq_of_dist_le_radius hp₁ hp₂.le).resolve_right h +end + +/-- Given three collinear points, two on a sphere and one not outside it, the one not outside it +is weakly between the other two points. -/ +lemma wbtw_of_collinear_of_dist_center_le_radius {s : sphere P} {p₁ p₂ p₃ : P} + (h : collinear ℝ ({p₁, p₂, p₃} : set P)) (hp₁ : p₁ ∈ s) (hp₂ : dist p₂ s.center ≤ s.radius) + (hp₃ : p₃ ∈ s) (hp₁p₃ : p₁ ≠ p₃) : wbtw ℝ p₁ p₂ p₃ := +h.wbtw_of_dist_eq_of_dist_le hp₁ hp₂ hp₃ hp₁p₃ + +/-- Given three collinear points, two on a sphere and one inside it, the one inside it is +strictly between the other two points. -/ +lemma sbtw_of_collinear_of_dist_center_lt_radius {s : sphere P} {p₁ p₂ p₃ : P} + (h : collinear ℝ ({p₁, p₂, p₃} : set P)) (hp₁ : p₁ ∈ s) (hp₂ : dist p₂ s.center < s.radius) + (hp₃ : p₃ ∈ s) (hp₁p₃ : p₁ ≠ p₃) : sbtw ℝ p₁ p₂ p₃ := +h.sbtw_of_dist_eq_of_dist_lt hp₁ hp₂ hp₃ hp₁p₃ + +/-- The second intersection of a sphere with a line through a point on that sphere; that point +if it is the only point of intersection of the line with the sphere. The intended use of this +definition is when `p ∈ s`; the definition does not use `s.radius`, so in general it returns +the second intersection with the sphere through `p` and with center `s.center`. -/ +def sphere.second_inter (s : sphere P) (p : P) (v : V) : P := +(-2 * ⟪v, p -ᵥ s.center⟫ / ⟪v, v⟫) • v +ᵥ p + +/-- The distance between `second_inter` and the center equals the distance between the original +point and the center. -/ +@[simp] lemma sphere.second_inter_dist (s : sphere P) (p : P) (v : V) : + dist (s.second_inter p v) s.center = dist p s.center := +begin + rw sphere.second_inter, + by_cases hv : v = 0, { simp [hv] }, + rw dist_smul_vadd_eq_dist _ _ hv, + exact or.inr rfl +end + +/-- The point given by `second_inter` lies on the sphere. -/ +@[simp] lemma sphere.second_inter_mem {s : sphere P} {p : P} (v : V) : + s.second_inter p v ∈ s ↔ p ∈ s := +by simp_rw [mem_sphere, sphere.second_inter_dist] + +variables (V) + +/-- If the vector is zero, `second_inter` gives the original point. -/ +@[simp] lemma sphere.second_inter_zero (s : sphere P) (p : P) : + s.second_inter p (0 : V) = p := +by simp [sphere.second_inter] + +variables {V} + +/-- The point given by `second_inter` equals the original point if and only if the line is +orthogonal to the radius vector. -/ +lemma sphere.second_inter_eq_self_iff {s : sphere P} {p : P} {v : V} : + s.second_inter p v = p ↔ ⟪v, p -ᵥ s.center⟫ = 0 := +begin + refine ⟨λ hp, _, λ hp, _⟩, + { by_cases hv : v = 0, { simp [hv] }, + rwa [sphere.second_inter, eq_comm, eq_vadd_iff_vsub_eq, vsub_self, eq_comm, smul_eq_zero, + or_iff_left hv, div_eq_zero_iff, inner_self_eq_zero, or_iff_left hv, mul_eq_zero, + or_iff_right (by norm_num : (-2 : ℝ) ≠ 0)] at hp }, + { rw [sphere.second_inter, hp, mul_zero, zero_div, zero_smul, zero_vadd] } +end + +/-- A point on a line through a point on a sphere equals that point or `second_inter`. -/ +lemma sphere.eq_or_eq_second_inter_of_mem_mk'_span_singleton_iff_mem {s : sphere P} {p : P} + (hp : p ∈ s) {v : V} {p' : P} (hp' : p' ∈ affine_subspace.mk' p (ℝ ∙ v)) : + (p' = p ∨ p' = s.second_inter p v) ↔ p' ∈ s := +begin + refine ⟨λ h, _, λ h, _⟩, + { rcases h with h | h, + { rwa h }, + { rwa [h, sphere.second_inter_mem] } }, + { rw [affine_subspace.mem_mk'_iff_vsub_mem, submodule.mem_span_singleton] at hp', + rcases hp' with ⟨r, hr⟩, + rw [eq_comm, ←eq_vadd_iff_vsub_eq] at hr, + subst hr, + by_cases hv : v = 0, { simp [hv] }, + rw sphere.second_inter, + rw mem_sphere at h hp, + rw [←hp, dist_smul_vadd_eq_dist _ _ hv] at h, + rcases h with h | h; + simp [h] } +end + +/-- `second_inter` is unchanged by multiplying the vector by a nonzero real. -/ +@[simp] lemma sphere.second_inter_smul (s : sphere P) (p : P) (v : V) {r : ℝ} + (hr : r ≠ 0) : s.second_inter p (r • v) = s.second_inter p v := +begin + simp_rw [sphere.second_inter, real_inner_smul_left, inner_smul_right, smul_smul, + div_mul_eq_div_div], + rw [mul_comm, ←mul_div_assoc, ←mul_div_assoc, mul_div_cancel_left _ hr, mul_comm, mul_assoc, + mul_div_cancel_left _ hr, mul_comm] +end + +/-- `second_inter` is unchanged by negating the vector. -/ +@[simp] lemma sphere.second_inter_neg (s : sphere P) (p : P) (v : V) : + s.second_inter p (-v) = s.second_inter p v := +by rw [←neg_one_smul ℝ v, s.second_inter_smul p v (by norm_num : (-1 : ℝ) ≠ 0)] + +/-- Applying `second_inter` twice returns the original point. -/ +@[simp] lemma sphere.second_inter_second_inter (s : sphere P) (p : P) (v : V) : + s.second_inter (s.second_inter p v) v = p := +begin + by_cases hv : v = 0, { simp [hv] }, + have hv' : ⟪v, v⟫ ≠ 0 := inner_self_eq_zero.not.2 hv, + simp only [sphere.second_inter, vadd_vsub_assoc, vadd_vadd, inner_add_right, inner_smul_right, + div_mul_cancel _ hv'], + rw [←@vsub_eq_zero_iff_eq V, vadd_vsub, ←add_smul, ←add_div], + convert zero_smul ℝ _, + convert zero_div _, + ring +end + +/-- If the vector passed to `second_inter` is given by a subtraction involving the point in +`second_inter`, the result of `second_inter` may be expressed using `line_map`. -/ +lemma sphere.second_inter_eq_line_map (s : sphere P) (p p' : P) : + s.second_inter p (p' -ᵥ p) = + affine_map.line_map p p' (-2 * ⟪p' -ᵥ p, p -ᵥ s.center⟫ / ⟪p' -ᵥ p, p' -ᵥ p⟫) := +rfl + +/-- If the vector passed to `second_inter` is given by a subtraction involving the point in +`second_inter`, the result lies in the span of the two points. -/ +lemma sphere.second_inter_vsub_mem_affine_span (s : sphere P) (p₁ p₂ : P) : + s.second_inter p₁ (p₂ -ᵥ p₁) ∈ line[ℝ, p₁, p₂] := +smul_vsub_vadd_mem_affine_span_pair _ _ _ + +/-- If the vector passed to `second_inter` is given by a subtraction involving the point in +`second_inter`, the three points are collinear. -/ +lemma sphere.second_inter_collinear (s : sphere P) (p p' : P) : + collinear ℝ ({p, p', s.second_inter p (p' -ᵥ p)} : set P) := +begin + rw [set.pair_comm, set.insert_comm], + exact (collinear_insert_iff_of_mem_affine_span (s.second_inter_vsub_mem_affine_span _ _)).2 + (collinear_pair ℝ _ _) +end + +/-- If the vector passed to `second_inter` is given by a subtraction involving the point in +`second_inter`, and the second point is not outside the sphere, the second point is weakly +between the first point and the result of `second_inter`. -/ +lemma sphere.wbtw_second_inter {s : sphere P} {p p' : P} (hp : p ∈ s) + (hp' : dist p' s.center ≤ s.radius) : wbtw ℝ p p' (s.second_inter p (p' -ᵥ p)) := +begin + by_cases h : p' = p, { simp [h] }, + refine wbtw_of_collinear_of_dist_center_le_radius (s.second_inter_collinear p p') + hp hp' ((sphere.second_inter_mem _).2 hp) _, + intro he, + rw [eq_comm, sphere.second_inter_eq_self_iff, ←neg_neg (p' -ᵥ p), inner_neg_left, + neg_vsub_eq_vsub_rev, neg_eq_zero, eq_comm] at he, + exact ((inner_pos_or_eq_of_dist_le_radius hp hp').resolve_right (ne.symm h)).ne he +end + +/-- If the vector passed to `second_inter` is given by a subtraction involving the point in +`second_inter`, and the second point is inside the sphere, the second point is strictly between +the first point and the result of `second_inter`. -/ +lemma sphere.sbtw_second_inter {s : sphere P} {p p' : P} (hp : p ∈ s) + (hp' : dist p' s.center < s.radius) : sbtw ℝ p p' (s.second_inter p (p' -ᵥ p)) := +begin + refine ⟨sphere.wbtw_second_inter hp hp'.le, _, _⟩, + { rintro rfl, rw mem_sphere at hp, simpa [hp] using hp' }, + { rintro h, + rw [h, mem_sphere.1 ((sphere.second_inter_mem _).2 hp)] at hp', + exact lt_irrefl _ hp' } +end + /-- A set of points is concyclic if it is cospherical and coplanar. (Most results are stated directly in terms of `cospherical` instead of using `concyclic`.) -/ structure concyclic (ps : set P) : Prop := From 0743cc5d9d86bcd1bba10f480e948a257d65056f Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Wed, 21 Dec 2022 07:58:40 +0000 Subject: [PATCH 0085/1291] chore(field_theory/separable_degree): golf a proof using `wlog` (#17952) --- src/field_theory/separable_degree.lean | 43 ++++++-------------------- 1 file changed, 10 insertions(+), 33 deletions(-) diff --git a/src/field_theory/separable_degree.lean b/src/field_theory/separable_degree.lean index c88eb9996e2c9..e91513df84c43 100644 --- a/src/field_theory/separable_degree.lean +++ b/src/field_theory/separable_degree.lean @@ -100,45 +100,22 @@ begin exact ⟨g, hgs, n, hge⟩, } end -/-- A helper lemma: if two expansions (along the positive characteristic) of two polynomials `g` and -`g'` agree, and the one with the larger degree is separable, then their degrees are the same. -/ -lemma contraction_degree_eq_aux [hq : fact q.prime] [hF : char_p F q] - (g g' : F[X]) (m m' : ℕ) - (h_expand : expand F (q^m) g = expand F (q^m') g') - (h : m < m') (hg : g.separable): - g.nat_degree = g'.nat_degree := -begin - obtain ⟨s, rfl⟩ := nat.exists_eq_add_of_lt h, - rw [add_assoc, pow_add, expand_mul] at h_expand, - let aux := expand_injective (pow_pos hq.1.pos m) h_expand, - rw aux at hg, - have := (is_unit_or_eq_zero_of_separable_expand q (s + 1) hq.out.pos hg).resolve_right - s.succ_ne_zero, - rw [aux, nat_degree_expand, - nat_degree_eq_of_degree_eq_some (degree_eq_zero_of_is_unit this), - zero_mul] -end - -/-- If two expansions (along the positive characteristic) of two separable polynomials -`g` and `g'` agree, then they have the same degree. -/ +/-- If two expansions (along the positive characteristic) of two separable polynomials `g` and `g'` +agree, then they have the same degree. -/ theorem contraction_degree_eq_or_insep - [hq : fact q.prime] [char_p F q] + [hq : ne_zero q] [char_p F q] (g g' : F[X]) (m m' : ℕ) (h_expand : expand F (q^m) g = expand F (q^m') g') (hg : g.separable) (hg' : g'.separable) : g.nat_degree = g'.nat_degree := begin - by_cases h : m = m', - { -- if `m = m'` then we show `g.nat_degree = g'.nat_degree` by unfolding the definitions - rw h at h_expand, - have expand_deg : ((expand F (q ^ m')) g).nat_degree = - (expand F (q ^ m') g').nat_degree, by rw h_expand, - rw [nat_degree_expand (q^m') g, nat_degree_expand (q^m') g'] at expand_deg, - apply nat.eq_of_mul_eq_mul_left (pow_pos hq.1.pos m'), - rw [mul_comm] at expand_deg, rw expand_deg, rw [mul_comm] }, - { cases ne.lt_or_lt h, - { exact contraction_degree_eq_aux q g g' m m' h_expand h_1 hg }, - { exact (contraction_degree_eq_aux q g' g m' m h_expand.symm h_1 hg').symm, } } + wlog hm : m ≤ m' := le_total m m' using [m m' g g', m' m g' g], + obtain ⟨s, rfl⟩ := exists_add_of_le hm, + rw [pow_add, expand_mul, expand_inj (pow_pos (ne_zero.pos q) m)] at h_expand, + subst h_expand, + rcases is_unit_or_eq_zero_of_separable_expand q s (ne_zero.pos q) hg with h | rfl, + { rw [nat_degree_expand, nat_degree_eq_zero_of_is_unit h, zero_mul] }, + { rw [nat_degree_expand, pow_zero, mul_one] }, end /-- The separable degree equals the degree of any separable contraction, i.e., it is unique. -/ From 6c48d300da03eaa8e374eb31933001090917202f Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Wed, 21 Dec 2022 11:23:29 +0000 Subject: [PATCH 0086/1291] feat(group_theory/subgroup/basic): add 2 lemmas (#17987) --- src/group_theory/subgroup/basic.lean | 5 +++++ 1 file changed, 5 insertions(+) diff --git a/src/group_theory/subgroup/basic.lean b/src/group_theory/subgroup/basic.lean index bf6bb5597f37e..2afc191ebafc8 100644 --- a/src/group_theory/subgroup/basic.lean +++ b/src/group_theory/subgroup/basic.lean @@ -2097,6 +2097,10 @@ lemma mem_ker (f : G →* M) {x : G} : x ∈ f.ker ↔ f x = 1 := iff.rfl @[to_additive] lemma coe_ker (f : G →* M) : (f.ker : set G) = (f : G → M) ⁻¹' {1} := rfl +@[simp, to_additive] +lemma ker_to_hom_units {M} [monoid M] (f : G →* M) : f.to_hom_units.ker = f.ker := +by { ext x, simp [mem_ker, units.ext_iff] } + @[to_additive] lemma eq_iff (f : G →* M) {x y : G} : f x = f y ↔ y⁻¹ * x ∈ f.ker := begin @@ -2127,6 +2131,7 @@ set_like.ext $ λ x, subtype.ext_iff ker_cod_restrict _ _ _ @[simp, to_additive] lemma ker_one : (1 : G →* M).ker = ⊤ := set_like.ext $ λ x, eq_self_iff_true _ +@[simp, to_additive] lemma ker_id : (monoid_hom.id G).ker = ⊥ := rfl @[to_additive] lemma ker_eq_bot_iff (f : G →* M) : f.ker = ⊥ ↔ function.injective f := ⟨λ h x y hxy, by rwa [eq_iff, h, mem_bot, inv_mul_eq_one, eq_comm] at hxy, From dce07aeb14157e75f17b69d6b816fa025e593765 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 21 Dec 2022 14:13:48 +0000 Subject: [PATCH 0087/1291] chore(data/real/cau_seq{_completion}): add lemmas about pow and smul (#17796) This also takes the opportunity to golf the algebraic instances, which should make porting easier. --- src/data/real/cau_seq.lean | 41 ++++++++++++--------- src/data/real/cau_seq_completion.lean | 53 ++++++++++++++------------- 2 files changed, 52 insertions(+), 42 deletions(-) diff --git a/src/data/real/cau_seq.lean b/src/data/real/cau_seq.lean index f988d639e7caa..1fd077feeb150 100644 --- a/src/data/real/cau_seq.lean +++ b/src/data/real/cau_seq.lean @@ -7,6 +7,8 @@ import algebra.group_power.lemmas import algebra.order.absolute_value import algebra.order.group.min_max import algebra.order.field.basic +import algebra.ring.pi +import group_theory.group_action.pi /-! # Cauchy sequences @@ -249,17 +251,14 @@ instance : has_smul G (cau_seq β abv) := @[simp, norm_cast] lemma smul_apply (a : G) (f : cau_seq β abv) (i : ℕ) : (a • f) i = a • f i := rfl lemma const_smul (a : G) (x : β) : const (a • x) = a • const x := rfl +instance : is_scalar_tower G (cau_seq β abv) (cau_seq β abv) := +⟨λ a f g, subtype.ext $ smul_assoc a ⇑f ⇑g⟩ + end has_smul instance : add_group (cau_seq β abv) := -by refine_struct - { add := (+), - neg := has_neg.neg, - zero := (0 : cau_seq β abv), - sub := has_sub.sub, - zsmul := (•), - nsmul := (•) }; -intros; try { refl }; apply ext; simp [add_comm, add_left_comm, sub_eq_add_neg, add_mul] +function.injective.add_group _ subtype.coe_injective + rfl coe_add coe_neg coe_sub (λ _ _, coe_smul _ _) (λ _ _, coe_smul _ _) instance : add_group_with_one (cau_seq β abv) := { one := 1, @@ -279,15 +278,9 @@ instance : has_pow (cau_seq β abv) ℕ := lemma const_pow (x : β) (n : ℕ) : const (x ^ n) = const x ^ n := rfl instance : ring (cau_seq β abv) := -by refine_struct - { add := (+), - zero := (0 : cau_seq β abv), - mul := (*), - one := 1, - npow := λ n f, f ^ n, - .. cau_seq.add_group_with_one }; -intros; try { refl }; apply ext; -simp [mul_add, mul_assoc, add_mul, add_comm, add_left_comm, sub_eq_add_neg, pow_succ] +function.injective.ring _ subtype.coe_injective + rfl rfl coe_add coe_mul coe_neg coe_sub (λ _ _, coe_smul _ _) (λ _ _, coe_smul _ _) coe_pow + (λ _, rfl) (λ _, rfl) instance {β : Type*} [comm_ring β] {abv : β → α} [is_absolute_value abv] : comm_ring (cau_seq β abv) := @@ -437,6 +430,20 @@ lemma mul_equiv_mul {f1 f2 g1 g2 : cau_seq β abv} (hf : f1 ≈ f2) (hg : g1 ≈ by simpa only [mul_sub, sub_mul, sub_add_sub_cancel] using add_lim_zero (mul_lim_zero_left g1 hf) (mul_lim_zero_right f2 hg) +lemma smul_equiv_smul [has_smul G β] [is_scalar_tower G β β] {f1 f2 : cau_seq β abv} + (c : G) (hf : f1 ≈ f2) : + c • f1 ≈ c • f2 := +by simpa [const_smul, smul_one_mul _ _] + using mul_equiv_mul (const_equiv.mpr $ eq.refl $ c • 1) hf + +lemma pow_equiv_pow {f1 f2 : cau_seq β abv} (hf : f1 ≈ f2) (n : ℕ) : + f1 ^ n ≈ f2 ^ n := +begin + induction n with n ih, + { simp only [pow_zero, setoid.refl] }, + { simpa only [pow_succ] using mul_equiv_mul hf ih, }, +end + end ring section is_domain diff --git a/src/data/real/cau_seq_completion.lean b/src/data/real/cau_seq_completion.lean index c71b1028d0bfb..40579d05b841d 100644 --- a/src/data/real/cau_seq_completion.lean +++ b/src/data/real/cau_seq_completion.lean @@ -63,6 +63,24 @@ instance : has_sub Cauchy := @[simp] theorem mk_sub (f g : cau_seq β abv) : mk f - mk g = mk (f - g) := rfl +instance {γ : Type*} [has_smul γ β] [is_scalar_tower γ β β] : has_smul γ Cauchy := +⟨λ c, quotient.map ((•) c) $ λ f₁ g₁ hf, smul_equiv_smul _ hf⟩ + +@[simp] theorem mk_smul {γ : Type*} [has_smul γ β] [is_scalar_tower γ β β] (c : γ) + (f : cau_seq β abv) : + c • mk f = mk (c • f) := rfl + +instance : has_pow Cauchy ℕ := +⟨λ x n, quotient.map (^ n) (λ f₁ g₁ hf, pow_equiv_pow hf _) x⟩ + +@[simp] theorem mk_pow (n : ℕ) (f : cau_seq β abv) : mk f ^ n = mk (f ^ n) := rfl + +instance : has_nat_cast Cauchy := ⟨λ n, mk n⟩ +instance : has_int_cast Cauchy := ⟨λ n, mk n⟩ + +@[simp] theorem of_rat_nat_cast (n : ℕ) : of_rat n = n := rfl +@[simp] theorem of_rat_int_cast (z : ℤ) : of_rat z = z := rfl + theorem of_rat_add (x y : β) : of_rat (x + y) = of_rat x + of_rat y := congr_arg mk (const_add _ _) @@ -76,30 +94,12 @@ private lemma zero_def : 0 = mk 0 := rfl private lemma one_def : 1 = mk 1 := rfl -instance : add_group Cauchy := -by refine { add := (+), zero := (0 : Cauchy), sub := has_sub.sub, neg := has_neg.neg, - sub_eq_add_neg := _, nsmul := nsmul_rec, zsmul := zsmul_rec, .. }; try { intros; refl }; -{ repeat {refine λ a, quotient.induction_on a (λ _, _)}, - simp [zero_def, add_comm, add_left_comm, sub_eq_neg_add] } - -instance : add_group_with_one Cauchy := -{ nat_cast := λ n, mk n, - nat_cast_zero := congr_arg mk nat.cast_zero, - nat_cast_succ := λ n, congr_arg mk (nat.cast_succ n), - int_cast := λ n, mk n, - int_cast_of_nat := λ n, congr_arg mk (int.cast_of_nat n), - int_cast_neg_succ_of_nat := λ n, congr_arg mk (int.cast_neg_succ_of_nat n), - one := 1, - .. Cauchy.add_group } - -@[simp] theorem of_rat_nat_cast (n : ℕ) : of_rat n = n := rfl -@[simp] theorem of_rat_int_cast (z : ℤ) : of_rat z = z := rfl - instance : ring Cauchy := -by refine { add := (+), zero := (0 : Cauchy), mul := (*), one := 1, npow := npow_rec, - .. Cauchy.add_group_with_one, .. }; try { intros; refl }; -{ repeat {refine λ a, quotient.induction_on a (λ _, _)}, - simp [zero_def, one_def, mul_add, add_mul, add_comm, add_left_comm, sub_eq_add_neg, ←mul_assoc] } +function.surjective.ring mk (surjective_quotient_mk _) + zero_def.symm one_def.symm (λ _ _, (mk_add _ _).symm) (λ _ _, (mk_mul _ _).symm) + (λ _, (mk_neg _).symm) (λ _ _, (mk_sub _ _).symm) + (λ _ _, (mk_smul _ _).symm) (λ _ _, (mk_smul _ _).symm) + (λ _ _, (mk_pow _ _).symm) (λ _, rfl) (λ _, rfl) /-- `cau_seq.completion.of_rat` as a `ring_hom` -/ @[simps] @@ -121,8 +121,11 @@ parameters {β : Type*} [comm_ring β] {abv : β → α} [is_absolute_value abv] local notation `Cauchy` := @Cauchy _ _ _ _ abv _ instance : comm_ring Cauchy := -{ mul_comm := quotient.ind₂ $ by exact λ a b, congr_arg quotient.mk $ mul_comm a b, - ..Cauchy.ring } +function.surjective.comm_ring mk (surjective_quotient_mk _) + zero_def.symm one_def.symm (λ _ _, (mk_add _ _).symm) (λ _ _, (mk_mul _ _).symm) + (λ _, (mk_neg _).symm) (λ _ _, (mk_sub _ _).symm) + (λ _ _, (mk_smul _ _).symm) (λ _ _, (mk_smul _ _).symm) + (λ _ _, (mk_pow _ _).symm) (λ _, rfl) (λ _, rfl) end From 3d95492390dc90e34184b13e865f50bc67f30fbb Mon Sep 17 00:00:00 2001 From: negiizhao Date: Wed, 21 Dec 2022 14:13:49 +0000 Subject: [PATCH 0088/1291] chore(data/set/lattice): move basic lemmas to basic files (#17882) This PR moves some lemmas in `data.set.lattice` to `data.set.basic`, `data.set.image`, and `data.set.function`. They are basic enough and do not depend on the set lattice. This also helps #17880 to split `data.set.pairwise` and reduce the dependency of many files. mathlib4 PR: https://github.com/leanprover-community/mathlib4/pull/1109 --- src/data/set/basic.lean | 156 +++++++++++++++++++++++++++ src/data/set/function.lean | 34 ++++++ src/data/set/image.lean | 60 ++++++++++- src/data/set/lattice.lean | 210 ------------------------------------- 4 files changed, 248 insertions(+), 212 deletions(-) diff --git a/src/data/set/basic.lean b/src/data/set/basic.lean index 97feb24603421..3f94e8ee93636 100644 --- a/src/data/set/basic.lean +++ b/src/data/set/basic.lean @@ -1736,3 +1736,159 @@ instance decidable_set_of (p : α → Prop) [decidable (p a)] : decidable (a ∈ by assumption end set + +/-! ### Monotone lemmas for sets -/ + +section monotone +variables {α β : Type*} + +theorem monotone.inter [preorder β] {f g : β → set α} + (hf : monotone f) (hg : monotone g) : monotone (λ x, f x ∩ g x) := +hf.inf hg + +theorem monotone_on.inter [preorder β] {f g : β → set α} {s : set β} + (hf : monotone_on f s) (hg : monotone_on g s) : monotone_on (λ x, f x ∩ g x) s := +hf.inf hg + +theorem antitone.inter [preorder β] {f g : β → set α} + (hf : antitone f) (hg : antitone g) : antitone (λ x, f x ∩ g x) := +hf.inf hg + +theorem antitone_on.inter [preorder β] {f g : β → set α} {s : set β} + (hf : antitone_on f s) (hg : antitone_on g s) : antitone_on (λ x, f x ∩ g x) s := +hf.inf hg + +theorem monotone.union [preorder β] {f g : β → set α} + (hf : monotone f) (hg : monotone g) : monotone (λ x, f x ∪ g x) := +hf.sup hg + +theorem monotone_on.union [preorder β] {f g : β → set α} {s : set β} + (hf : monotone_on f s) (hg : monotone_on g s) : monotone_on (λ x, f x ∪ g x) s := +hf.sup hg + +theorem antitone.union [preorder β] {f g : β → set α} + (hf : antitone f) (hg : antitone g) : antitone (λ x, f x ∪ g x) := +hf.sup hg + +theorem antitone_on.union [preorder β] {f g : β → set α} {s : set β} + (hf : antitone_on f s) (hg : antitone_on g s) : antitone_on (λ x, f x ∪ g x) s := +hf.sup hg + +namespace set + +theorem monotone_set_of [preorder α] {p : α → β → Prop} + (hp : ∀ b, monotone (λ a, p a b)) : monotone (λ a, {b | p a b}) := +λ a a' h b, hp b h + +theorem antitone_set_of [preorder α] {p : α → β → Prop} + (hp : ∀ b, antitone (λ a, p a b)) : antitone (λ a, {b | p a b}) := +λ a a' h b, hp b h + +/-- Quantifying over a set is antitone in the set -/ +lemma antitone_bforall {P : α → Prop} : antitone (λ s : set α, ∀ x ∈ s, P x) := +λ s t hst h x hx, h x $ hst hx + +end set + +end monotone + +/-! ### Disjoint sets -/ + +section disjoint + +variables {α β : Type*} {s t u : set α} {f : α → β} + +namespace disjoint + +theorem union_left (hs : disjoint s u) (ht : disjoint t u) : disjoint (s ∪ t) u := +hs.sup_left ht + +theorem union_right (ht : disjoint s t) (hu : disjoint s u) : disjoint s (t ∪ u) := +ht.sup_right hu + +lemma inter_left (u : set α) (h : disjoint s t) : disjoint (s ∩ u) t := +h.inf_left u + +lemma inter_left' (u : set α) (h : disjoint s t) : disjoint (u ∩ s) t := +h.inf_left' _ + +lemma inter_right (u : set α) (h : disjoint s t) : disjoint s (t ∩ u) := +h.inf_right _ + +lemma inter_right' (u : set α) (h : disjoint s t) : disjoint s (u ∩ t) := +h.inf_right' _ + +lemma subset_left_of_subset_union (h : s ⊆ t ∪ u) (hac : disjoint s u) : s ⊆ t := +hac.left_le_of_le_sup_right h + +lemma subset_right_of_subset_union (h : s ⊆ t ∪ u) (hab : disjoint s t) : s ⊆ u := +hab.left_le_of_le_sup_left h + +end disjoint + +namespace set + +lemma not_disjoint_iff : ¬disjoint s t ↔ ∃ x, x ∈ s ∧ x ∈ t := +set.disjoint_iff.not.trans $ not_forall.trans $ exists_congr $ λ x, not_not + +lemma not_disjoint_iff_nonempty_inter : ¬disjoint s t ↔ (s ∩ t).nonempty := +not_disjoint_iff + +alias not_disjoint_iff_nonempty_inter ↔ _ nonempty.not_disjoint + +lemma disjoint_or_nonempty_inter (s t : set α) : disjoint s t ∨ (s ∩ t).nonempty := +(em _).imp_right not_disjoint_iff_nonempty_inter.mp + +lemma disjoint_iff_forall_ne : disjoint s t ↔ ∀ (x ∈ s) (y ∈ t), x ≠ y := +by simp only [ne.def, disjoint_left, @imp_not_comm _ (_ = _), forall_eq'] + +lemma _root_.disjoint.ne_of_mem (h : disjoint s t) {x y} (hx : x ∈ s) (hy : y ∈ t) : x ≠ y := +disjoint_iff_forall_ne.mp h x hx y hy + +theorem disjoint_of_subset_left (h : s ⊆ u) (d : disjoint u t) : disjoint s t := +d.mono_left h + +theorem disjoint_of_subset_right (h : t ⊆ u) (d : disjoint s u) : disjoint s t := +d.mono_right h + +theorem disjoint_of_subset {s t u v : set α} (h1 : s ⊆ u) (h2 : t ⊆ v) (d : disjoint u v) : + disjoint s t := +d.mono h1 h2 + +@[simp] theorem disjoint_union_left : + disjoint (s ∪ t) u ↔ disjoint s u ∧ disjoint t u := +disjoint_sup_left + +@[simp] theorem disjoint_union_right : + disjoint s (t ∪ u) ↔ disjoint s t ∧ disjoint s u := +disjoint_sup_right + +theorem disjoint_diff {a b : set α} : disjoint a (b \ a) := +disjoint_iff.2 (inter_diff_self _ _) + +@[simp] theorem disjoint_empty (s : set α) : disjoint s ∅ := disjoint_bot_right + +@[simp] theorem empty_disjoint (s : set α) : disjoint ∅ s := disjoint_bot_left + +@[simp] lemma univ_disjoint {s : set α} : disjoint univ s ↔ s = ∅ := +top_disjoint + +@[simp] lemma disjoint_univ {s : set α} : disjoint s univ ↔ s = ∅ := +disjoint_top + +@[simp] theorem disjoint_singleton_left {a : α} {s : set α} : disjoint {a} s ↔ a ∉ s := +by simp [set.disjoint_iff, subset_def]; exact iff.rfl + +@[simp] theorem disjoint_singleton_right {a : α} {s : set α} : disjoint s {a} ↔ a ∉ s := +by rw [disjoint.comm]; exact disjoint_singleton_left + +@[simp] lemma disjoint_singleton {a b : α} : disjoint ({a} : set α) {b} ↔ a ≠ b := +by rw [disjoint_singleton_left, mem_singleton_iff] + +lemma subset_diff {s t u : set α} : s ⊆ t \ u ↔ s ⊆ t ∧ disjoint s u := +⟨λ h, ⟨λ x hxs, (h hxs).1, disjoint_iff_inf_le.mpr $ λ x ⟨hxs, hxu⟩, (h hxs).2 hxu⟩, +λ ⟨h1, h2⟩ x hxs, ⟨h1 hxs, λ hxu, h2.le_bot ⟨hxs, hxu⟩⟩⟩ + +end set + +end disjoint diff --git a/src/data/set/function.lean b/src/data/set/function.lean index 60e82ebc9c11b..07be41cf6a539 100644 --- a/src/data/set/function.lean +++ b/src/data/set/function.lean @@ -411,6 +411,21 @@ begin set.preimage_inter, subtype.coe_preimage_self, set.univ_inter], end +variables {f} {U : ι → set β} + +lemma restrict_preimage_injective (hf : injective f) : injective (t.restrict_preimage f) := +λ x y e, subtype.mk.inj_arrow e (λ e, subtype.coe_injective (hf e)) + +lemma restrict_preimage_surjective (hf : surjective f) : surjective (t.restrict_preimage f) := +λ x, ⟨⟨_, (show f (hf x).some ∈ t, from (hf x).some_spec.symm ▸ x.2)⟩, subtype.ext (hf x).some_spec⟩ + +lemma restrict_preimage_bijective (hf : bijective f) : bijective (t.restrict_preimage f) := +⟨t.restrict_preimage_injective hf.1, t.restrict_preimage_surjective hf.2⟩ + +alias set.restrict_preimage_injective ← _root_.function.injective.restrict_preimage +alias set.restrict_preimage_surjective ← _root_.function.surjective.restrict_preimage +alias set.restrict_preimage_bijective ← _root_.function.bijective.restrict_preimage + end /-! ### Injectivity on a set -/ @@ -518,6 +533,25 @@ lemma inj_on.cancel_left (hg : t.inj_on g) (hf₁ : s.maps_to f₁ t) (hf₂ : s s.eq_on (g ∘ f₁) (g ∘ f₂) ↔ s.eq_on f₁ f₂ := ⟨λ h, h.cancel_left hg hf₁ hf₂, eq_on.comp_left⟩ +lemma inj_on.image_inter {s t u : set α} (hf : u.inj_on f) (hs : s ⊆ u) (ht : t ⊆ u) : + f '' (s ∩ t) = f '' s ∩ f '' t := +begin + apply subset.antisymm (image_inter_subset _ _ _), + rintros x ⟨⟨y, ys, hy⟩, ⟨z, zt, hz⟩⟩, + have : y = z, + { apply hf (hs ys) (ht zt), + rwa ← hz at hy }, + rw ← this at zt, + exact ⟨y, ⟨ys, zt⟩, hy⟩, +end + +lemma _root_.disjoint.image {s t u : set α} {f : α → β} (h : disjoint s t) (hf : inj_on f u) + (hs : s ⊆ u) (ht : t ⊆ u) : disjoint (f '' s) (f '' t) := +begin + rw disjoint_iff_inter_eq_empty at h ⊢, + rw [← hf.image_inter hs ht, h, image_empty], +end + /-! ### Surjectivity on a set -/ /-- `f` is surjective from `a` to `b` if `b` is contained in the image of `a`. -/ diff --git a/src/data/set/image.lean b/src/data/set/image.lean index 0a6463c1cb78d..f1606759efb09 100644 --- a/src/data/set/image.lean +++ b/src/data/set/image.lean @@ -134,7 +134,6 @@ end end preimage - /-! ### Image of a set under a function -/ section image @@ -217,6 +216,10 @@ terms of `≤`. -/ theorem image_subset {a b : set α} (f : α → β) (h : a ⊆ b) : f '' a ⊆ f '' b := by { simp only [subset_def, mem_image], exact λ x, λ ⟨w, h1, h2⟩, ⟨w, h h1, h2⟩ } +/-- `set.image` is monotone. See `set.image_subset` for the statement in terms of `⊆`. -/ +lemma monotone_image {f : α → β} : monotone (image f) := +λ s t, image_subset _ + theorem image_union (f : α → β) (s t : set α) : f '' (s ∪ t) = f '' s ∪ f '' t := ext $ λ x, ⟨by rintro ⟨a, h|h, rfl⟩; [left, right]; exact ⟨_, h, rfl⟩, @@ -1101,7 +1104,7 @@ option.range_eq f namespace set open function -/-! ### Injectivity and sur Date: Wed, 21 Dec 2022 14:13:50 +0000 Subject: [PATCH 0089/1291] feat(*): add various results split out from proof of Gallagher's theorem (#17985) This is a collection of tiny results that are not really related but I'm bundling them together in an effort to aid review. --- src/analysis/normed/field/basic.lean | 37 +++++------------- src/analysis/normed/group/basic.lean | 36 ++++++++++++++++++ src/analysis/normed_space/int.lean | 2 - src/data/nat/totient.lean | 11 ++++++ .../covering/liminf_limsup.lean | 2 +- src/measure_theory/measurable_space.lean | 26 +++++++++++-- .../measure/measure_space_def.lean | 32 ++++++++++++++++ src/order/hom/complete_lattice.lean | 31 ++++++++++++++- src/order/liminf_limsup.lean | 38 ++++++++++++++++--- 9 files changed, 174 insertions(+), 41 deletions(-) diff --git a/src/analysis/normed/field/basic.lean b/src/analysis/normed/field/basic.lean index ea3f874e13211..15dd3438859d1 100644 --- a/src/analysis/normed/field/basic.lean +++ b/src/analysis/normed/field/basic.lean @@ -693,6 +693,8 @@ instance : normed_comm_ring ℤ := lemma int.norm_eq_abs (n : ℤ) : ‖n‖ = |n| := rfl +@[simp] lemma int.norm_coe_nat (n : ℕ) : ‖(n : ℤ)‖ = n := by simp [int.norm_eq_abs] + lemma nnreal.coe_nat_abs (n : ℤ) : (n.nat_abs : ℝ≥0) = ‖n‖₊ := nnreal.eq $ calc ((n.nat_abs : ℝ≥0) : ℝ) = (n.nat_abs : ℤ) : by simp only [int.cast_coe_nat, nnreal.coe_nat_cast] @@ -724,36 +726,17 @@ instance : densely_normed_field ℚ := by rw [← rat.norm_cast_real, ← int.norm_cast_real]; congr' 1; norm_cast -- Now that we've installed the norm on `ℤ`, --- we can state some lemmas about `nsmul` and `zsmul`. +-- we can state some lemmas about `zsmul`. section -variables [seminormed_add_comm_group α] - -lemma norm_nsmul_le (n : ℕ) (a : α) : ‖n • a‖ ≤ n * ‖a‖ := -begin - induction n with n ih, - { simp only [norm_zero, nat.cast_zero, zero_mul, zero_smul] }, - simp only [nat.succ_eq_add_one, add_smul, add_mul, one_mul, nat.cast_add, - nat.cast_one, one_nsmul], - exact norm_add_le_of_le ih le_rfl -end - -lemma norm_zsmul_le (n : ℤ) (a : α) : ‖n • a‖ ≤ ‖n‖ * ‖a‖ := -begin - induction n with n n, - { simp only [int.of_nat_eq_coe, coe_nat_zsmul], - convert norm_nsmul_le n a, - exact nat.abs_cast n }, - { simp only [int.neg_succ_of_nat_coe, neg_smul, norm_neg, coe_nat_zsmul], - convert norm_nsmul_le n.succ a, - exact nat.abs_cast n.succ, } -end +variables [seminormed_comm_group α] -lemma nnnorm_nsmul_le (n : ℕ) (a : α) : ‖n • a‖₊ ≤ n * ‖a‖₊ := -by simpa only [←nnreal.coe_le_coe, nnreal.coe_mul, nnreal.coe_nat_cast] - using norm_nsmul_le n a +@[to_additive norm_zsmul_le] +lemma norm_zpow_le_mul_norm (n : ℤ) (a : α) : ‖a^n‖ ≤ ‖n‖ * ‖a‖ := +by rcases n.eq_coe_or_neg with ⟨n, rfl | rfl⟩; simpa using norm_pow_le_mul_norm n a -lemma nnnorm_zsmul_le (n : ℤ) (a : α) : ‖n • a‖₊ ≤ ‖n‖₊ * ‖a‖₊ := -by simpa only [←nnreal.coe_le_coe, nnreal.coe_mul] using norm_zsmul_le n a +@[to_additive nnnorm_zsmul_le] +lemma nnnorm_zpow_le_mul_norm (n : ℤ) (a : α) : ‖a^n‖₊ ≤ ‖n‖₊ * ‖a‖₊ := +by simpa only [← nnreal.coe_le_coe, nnreal.coe_mul] using norm_zpow_le_mul_norm n a end diff --git a/src/analysis/normed/group/basic.lean b/src/analysis/normed/group/basic.lean index abbeb4e1b952e..fae033dd6ec4c 100644 --- a/src/analysis/normed/group/basic.lean +++ b/src/analysis/normed/group/basic.lean @@ -1061,6 +1061,42 @@ by { ext c, ((*) b) ⁻¹' sphere a r = sphere (a / b) r := by { ext c, simp only [set.mem_preimage, mem_sphere_iff_norm', div_div_eq_mul_div, mul_comm] } +@[to_additive norm_nsmul_le] lemma norm_pow_le_mul_norm (n : ℕ) (a : E) : ‖a^n‖ ≤ n * ‖a‖ := +begin + induction n with n ih, { simp, }, + simpa only [pow_succ', nat.cast_succ, add_mul, one_mul] using norm_mul_le_of_le ih le_rfl, +end + +@[to_additive nnnorm_nsmul_le] lemma nnnorm_pow_le_mul_norm (n : ℕ) (a : E) : ‖a^n‖₊ ≤ n * ‖a‖₊ := +by simpa only [← nnreal.coe_le_coe, nnreal.coe_mul, nnreal.coe_nat_cast] + using norm_pow_le_mul_norm n a + +@[to_additive] lemma pow_mem_closed_ball {n : ℕ} (h : a ∈ closed_ball b r) : + a^n ∈ closed_ball (b^n) (n • r) := +begin + simp only [mem_closed_ball, dist_eq_norm_div, ← div_pow] at h ⊢, + refine (norm_pow_le_mul_norm n (a / b)).trans _, + simpa only [nsmul_eq_mul] using mul_le_mul_of_nonneg_left h n.cast_nonneg, +end + +@[to_additive] lemma pow_mem_ball {n : ℕ} (hn : 0 < n) (h : a ∈ ball b r) : + a^n ∈ ball (b^n) (n • r) := +begin + simp only [mem_ball, dist_eq_norm_div, ← div_pow] at h ⊢, + refine lt_of_le_of_lt (norm_pow_le_mul_norm n (a / b)) _, + replace hn : 0 < (n : ℝ), { norm_cast, assumption, }, + rw nsmul_eq_mul, + nlinarith, +end + +@[simp, to_additive] lemma mul_mem_closed_ball_mul_iff {c : E} : + a * c ∈ closed_ball (b * c) r ↔ a ∈ closed_ball b r := +by simp only [mem_closed_ball, dist_eq_norm_div, mul_div_mul_right_eq_div] + +@[simp, to_additive] lemma mul_mem_ball_mul_iff {c : E} : + a * c ∈ ball (b * c) r ↔ a ∈ ball b r := +by simp only [mem_ball, dist_eq_norm_div, mul_div_mul_right_eq_div] + namespace isometric /-- Multiplication `y ↦ x * y` as an `isometry`. -/ diff --git a/src/analysis/normed_space/int.lean b/src/analysis/normed_space/int.lean index 786f3f7d5fe07..51d30e87fa72d 100644 --- a/src/analysis/normed_space/int.lean +++ b/src/analysis/normed_space/int.lean @@ -31,8 +31,6 @@ by rw [← coe_nnnorm, int.nnnorm_coe_units, nnreal.coe_one] @[simp] lemma nnnorm_coe_nat (n : ℕ) : ‖(n : ℤ)‖₊ = n := real.nnnorm_coe_nat _ -@[simp] lemma norm_coe_nat (n : ℕ) : ‖(n : ℤ)‖ = n := real.norm_coe_nat _ - @[simp] lemma to_nat_add_to_nat_neg_eq_nnnorm (n : ℤ) : ↑(n.to_nat) + ↑((-n).to_nat) = ‖n‖₊ := by rw [← nat.cast_add, to_nat_add_to_nat_neg_eq_nat_abs, nnreal.coe_nat_abs] diff --git a/src/data/nat/totient.lean b/src/data/nat/totient.lean index d78eae8ae6b7c..b3082bdf5a895 100644 --- a/src/data/nat/totient.lean +++ b/src/data/nat/totient.lean @@ -36,6 +36,17 @@ by simp [totient] lemma totient_eq_card_coprime (n : ℕ) : φ n = ((range n).filter n.coprime).card := rfl +/-- A characterisation of `nat.totient` that avoids `finset`. -/ +lemma totient_eq_card_lt_and_coprime (n : ℕ) : φ n = nat.card {m | m < n ∧ n.coprime m} := +begin + let e : {m | m < n ∧ n.coprime m} ≃ finset.filter n.coprime (finset.range n) := + { to_fun := λ m, ⟨m, by simpa only [finset.mem_filter, finset.mem_range] using m.property⟩, + inv_fun := λ m, ⟨m, by simpa only [finset.mem_filter, finset.mem_range] using m.property⟩, + left_inv := λ m, by simp only [subtype.coe_mk, subtype.coe_eta], + right_inv := λ m, by simp only [subtype.coe_mk, subtype.coe_eta] }, + rw [totient_eq_card_coprime, card_congr e, card_eq_fintype_card, fintype.card_coe], +end + lemma totient_le (n : ℕ) : φ n ≤ n := ((range n).card_filter_le _).trans_eq (card_range n) diff --git a/src/measure_theory/covering/liminf_limsup.lean b/src/measure_theory/covering/liminf_limsup.lean index 11eb21e62c41e..bf38500e1ec2d 100644 --- a/src/measure_theory/covering/liminf_limsup.lean +++ b/src/measure_theory/covering/liminf_limsup.lean @@ -163,7 +163,7 @@ begin cases le_or_lt 1 M with hM' hM', { apply has_subset.subset.eventually_le, change _ ≤ _, - refine mono_blimsup' (hMr.mono $ λ i hi, cthickening_mono _ (s i)), + refine mono_blimsup' (hMr.mono $ λ i hi hp, cthickening_mono _ (s i)), exact (le_mul_of_one_le_left (hRp i) hM').trans hi, }, { simp only [← @cthickening_closure _ _ _ (s _)], have hs : ∀ i, is_closed (closure (s i)) := λ i, is_closed_closure, diff --git a/src/measure_theory/measurable_space.lean b/src/measure_theory/measurable_space.lean index 790360725b97b..821af074c9ce5 100644 --- a/src/measure_theory/measurable_space.lean +++ b/src/measure_theory/measurable_space.lean @@ -1458,18 +1458,36 @@ instance : boolean_algebra (subtype (measurable_set : set α → Prop)) := .. measurable_set.subtype.bounded_order, .. measurable_set.subtype.distrib_lattice } +@[measurability] lemma measurable_set_blimsup {s : ℕ → set α} {p : ℕ → Prop} + (h : ∀ n, p n → measurable_set (s n)) : + measurable_set $ filter.blimsup s filter.at_top p := +begin + simp only [filter.blimsup_eq_infi_bsupr_of_nat, supr_eq_Union, infi_eq_Inter], + exact measurable_set.Inter + (λ n, measurable_set.Union (λ m, measurable_set.Union $ λ hm, h m hm.1)), +end + +@[measurability] lemma measurable_set_bliminf {s : ℕ → set α} {p : ℕ → Prop} + (h : ∀ n, p n → measurable_set (s n)) : + measurable_set $ filter.bliminf s filter.at_top p := +begin + simp only [filter.bliminf_eq_supr_binfi_of_nat, infi_eq_Inter, supr_eq_Union], + exact measurable_set.Union + (λ n, measurable_set.Inter (λ m, measurable_set.Inter $ λ hm, h m hm.1)), +end + @[measurability] lemma measurable_set_limsup {s : ℕ → set α} (hs : ∀ n, measurable_set $ s n) : measurable_set $ filter.limsup s filter.at_top := begin - simp only [filter.limsup_eq_infi_supr_of_nat', supr_eq_Union, infi_eq_Inter], - exact measurable_set.Inter (λ n, measurable_set.Union $ λ m, hs $ m + n), + convert measurable_set_blimsup (λ n h, hs n : ∀ n, true → measurable_set (s n)), + simp, end @[measurability] lemma measurable_set_liminf {s : ℕ → set α} (hs : ∀ n, measurable_set $ s n) : measurable_set $ filter.liminf s filter.at_top := begin - simp only [filter.liminf_eq_supr_infi_of_nat', supr_eq_Union, infi_eq_Inter], - exact measurable_set.Union (λ n, measurable_set.Inter $ λ m, hs $ m + n), + convert measurable_set_bliminf (λ n h, hs n : ∀ n, true → measurable_set (s n)), + simp, end end measurable_set diff --git a/src/measure_theory/measure/measure_space_def.lean b/src/measure_theory/measure/measure_space_def.lean index 06ca7e30abe12..1f61e6e5d47b5 100644 --- a/src/measure_theory/measure/measure_space_def.lean +++ b/src/measure_theory/measure/measure_space_def.lean @@ -406,6 +406,38 @@ lemma ae_eq_set_union {s' t' : set α} (h : s =ᵐ[μ] t) (h' : s' =ᵐ[μ] t') (s ∪ s' : set α) =ᵐ[μ] (t ∪ t' : set α) := h.union h' +lemma union_ae_eq_univ_of_ae_eq_univ_left (h : s =ᵐ[μ] univ) : + (s ∪ t : set α) =ᵐ[μ] univ := +by { convert ae_eq_set_union h (ae_eq_refl t), rw univ_union, } + +lemma union_ae_eq_univ_of_ae_eq_univ_right (h : t =ᵐ[μ] univ) : + (s ∪ t : set α) =ᵐ[μ] univ := +by { convert ae_eq_set_union (ae_eq_refl s) h, rw union_univ, } + +lemma union_ae_eq_right_of_ae_eq_empty (h : s =ᵐ[μ] (∅ : set α)) : + (s ∪ t : set α) =ᵐ[μ] t := +by { convert ae_eq_set_union h (ae_eq_refl t), rw empty_union, } + +lemma union_ae_eq_left_of_ae_eq_empty (h : t =ᵐ[μ] (∅ : set α)) : + (s ∪ t : set α) =ᵐ[μ] s := +by { convert ae_eq_set_union (ae_eq_refl s) h, rw union_empty, } + +lemma inter_ae_eq_right_of_ae_eq_univ (h : s =ᵐ[μ] univ) : + (s ∩ t : set α) =ᵐ[μ] t := +by { convert ae_eq_set_inter h (ae_eq_refl t), rw univ_inter, } + +lemma inter_ae_eq_left_of_ae_eq_univ (h : t =ᵐ[μ] univ) : + (s ∩ t : set α) =ᵐ[μ] s := +by { convert ae_eq_set_inter (ae_eq_refl s) h, rw inter_univ, } + +lemma inter_ae_eq_empty_of_ae_eq_empty_left (h : s =ᵐ[μ] (∅ : set α)) : + (s ∩ t : set α) =ᵐ[μ] (∅ : set α) := +by { convert ae_eq_set_inter h (ae_eq_refl t), rw empty_inter, } + +lemma inter_ae_eq_empty_of_ae_eq_empty_right (h : t =ᵐ[μ] (∅ : set α)) : + (s ∩ t : set α) =ᵐ[μ] (∅ : set α) := +by { convert ae_eq_set_inter (ae_eq_refl s) h, rw inter_empty, } + @[to_additive] lemma _root_.set.mul_indicator_ae_eq_one {M : Type*} [has_one M] {f : α → M} {s : set α} (h : s.mul_indicator f =ᵐ[μ] 1) : μ (s ∩ function.mul_support f) = 0 := diff --git a/src/order/hom/complete_lattice.lean b/src/order/hom/complete_lattice.lean index d40a85a082fe9..7d3b73ead69ac 100644 --- a/src/order/hom/complete_lattice.lean +++ b/src/order/hom/complete_lattice.lean @@ -583,7 +583,9 @@ end complete_lattice_hom namespace complete_lattice_hom -/-- `set.preimage` as a complete lattice homomorphism. -/ +/-- `set.preimage` as a complete lattice homomorphism. + +See also `Sup_hom.set_image`. -/ def set_preimage (f : α → β) : complete_lattice_hom (set β) (set α) := { to_fun := preimage f, map_Sup' := λ s, preimage_sUnion.trans $ by simp only [set.Sup_eq_sUnion, set.sUnion_image], @@ -597,3 +599,30 @@ lemma set_preimage_comp (g : β → γ) (f : α → β) : set_preimage (g ∘ f) = (set_preimage f).comp (set_preimage g) := rfl end complete_lattice_hom + +lemma set.image_Sup {f : α → β} (s : set (set α)) : + f '' Sup s = Sup (image f '' s) := +begin + ext b, + simp only [Sup_eq_sUnion, mem_image, mem_sUnion, exists_prop, sUnion_image, mem_Union], + split, + { rintros ⟨a, ⟨t, ht₁, ht₂⟩, rfl⟩, exact ⟨t, ht₁, a, ht₂, rfl⟩, }, + { rintros ⟨t, ht₁, a, ht₂, rfl⟩, exact ⟨a, ⟨t, ht₁, ht₂⟩, rfl⟩, }, +end + +/-- Using `set.image`, a function between types yields a `Sup_hom` between their lattices of +subsets. + +See also `complete_lattice_hom.set_preimage`. -/ +@[simps] def Sup_hom.set_image (f : α → β) : Sup_hom (set α) (set β) := +{ to_fun := image f, + map_Sup' := set.image_Sup } + +/-- An equivalence of types yields an order isomorphism between their lattices of subsets. -/ +@[simps] def equiv.to_order_iso_set (e : α ≃ β) : set α ≃o set β := +{ to_fun := image e, + inv_fun := image e.symm, + left_inv := λ s, by simp only [← image_comp, equiv.symm_comp_self, id.def, image_id'], + right_inv := λ s, by simp only [← image_comp, equiv.self_comp_symm, id.def, image_id'], + map_rel_iff' := + λ s t, ⟨λ h, by simpa using @monotone_image _ _ e.symm _ _ h, λ h, monotone_image h⟩ } diff --git a/src/order/liminf_limsup.lean b/src/order/liminf_limsup.lean index d0200be270336..effad5afac2bb 100644 --- a/src/order/liminf_limsup.lean +++ b/src/order/liminf_limsup.lean @@ -661,19 +661,19 @@ lemma bliminf_antitone (h : ∀ x, p x → q x) : bliminf u f q ≤ bliminf u f p := Sup_le_Sup $ λ a ha, ha.mono $ by tauto -lemma mono_blimsup' (h : ∀ᶠ x in f, u x ≤ v x) : +lemma mono_blimsup' (h : ∀ᶠ x in f, p x → u x ≤ v x) : blimsup u f p ≤ blimsup v f p := -Inf_le_Inf $ λ a ha, (ha.and h).mono $ λ x hx hx', hx.2.trans (hx.1 hx') +Inf_le_Inf $ λ a ha, (ha.and h).mono $ λ x hx hx', (hx.2 hx').trans (hx.1 hx') -lemma mono_blimsup (h : ∀ x, u x ≤ v x) : +lemma mono_blimsup (h : ∀ x, p x → u x ≤ v x) : blimsup u f p ≤ blimsup v f p := mono_blimsup' $ eventually_of_forall h -lemma mono_bliminf' (h : ∀ᶠ x in f, u x ≤ v x) : +lemma mono_bliminf' (h : ∀ᶠ x in f, p x → u x ≤ v x) : bliminf u f p ≤ bliminf v f p := -Sup_le_Sup $ λ a ha, (ha.and h).mono $ λ x hx hx', (hx.1 hx').trans hx.2 +Sup_le_Sup $ λ a ha, (ha.and h).mono $ λ x hx hx', (hx.1 hx').trans (hx.2 hx') -lemma mono_bliminf (h : ∀ x, u x ≤ v x) : +lemma mono_bliminf (h : ∀ x, p x → u x ≤ v x) : bliminf u f p ≤ bliminf v f p := mono_bliminf' $ eventually_of_forall h @@ -703,6 +703,32 @@ sup_le (blimsup_mono $ by tauto) (blimsup_mono $ by tauto) bliminf u f (λ x, p x ∨ q x) ≤ bliminf u f p ⊓ bliminf u f q := @blimsup_sup_le_or αᵒᵈ β _ f p q u +lemma order_iso.apply_blimsup [complete_lattice γ] (e : α ≃o γ) : + e (blimsup u f p) = blimsup (e ∘ u) f p := +begin + simp only [blimsup_eq, map_Inf, function.comp_app], + congr, + ext c, + obtain ⟨a, rfl⟩ := e.surjective c, + simp, +end + +lemma order_iso.apply_bliminf [complete_lattice γ] (e : α ≃o γ) : + e (bliminf u f p) = bliminf (e ∘ u) f p := +@order_iso.apply_blimsup αᵒᵈ β γᵒᵈ _ f p u _ e.dual + +lemma Sup_hom.apply_blimsup_le [complete_lattice γ] (g : Sup_hom α γ) : + g (blimsup u f p) ≤ blimsup (g ∘ u) f p := +begin + simp only [blimsup_eq_infi_bsupr], + refine ((order_hom_class.mono g).map_infi₂_le _).trans _, + simp only [_root_.map_supr], +end + +lemma Inf_hom.le_apply_bliminf [complete_lattice γ] (g : Inf_hom α γ) : + bliminf (g ∘ u) f p ≤ g (bliminf u f p) := +@Sup_hom.apply_blimsup_le αᵒᵈ β γᵒᵈ _ f p u _ g.dual + end complete_lattice section complete_distrib_lattice From 56f4cd1ef396e9fd389b5d8371ee9ad91d163625 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Wed, 21 Dec 2022 17:20:44 +0000 Subject: [PATCH 0090/1291] feat(measure_theory/kernel/basic): define kernels and classes of kernels (Markov, finite, s-finite) (#17976) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We define kernels between measurable spaces as the `add_submonoid` of measurable functions `α → measure β`. That is, a kernel `κ` verifies that for all measurable sets `s` of `β`, `λ a, κ a s` is measurable. A kernel is said to be Markov if all measures its image are probability measures, finite if all measures are finite with a common bound, and s-finite if it can be written as a countable sum of finite kernels. --- src/measure_theory/integral/lebesgue.lean | 10 + src/measure_theory/measure/giry_monad.lean | 9 + src/measure_theory/measure/measure_space.lean | 4 + src/probability/kernel/basic.lean | 294 ++++++++++++++++++ 4 files changed, 317 insertions(+) create mode 100644 src/probability/kernel/basic.lean diff --git a/src/measure_theory/integral/lebesgue.lean b/src/measure_theory/integral/lebesgue.lean index e777ec419eb0a..fa276733d378f 100644 --- a/src/measure_theory/integral/lebesgue.lean +++ b/src/measure_theory/integral/lebesgue.lean @@ -1662,6 +1662,10 @@ by rw [← lintegral_congr_ae (indicator_ae_eq_of_ae_eq_set hs.to_measurable_ae_ lintegral_indicator _ (measurable_set_to_measurable _ _), measure.restrict_congr_set hs.to_measurable_ae_eq] +lemma lintegral_indicator_const {s : set α} (hs : measurable_set s) (c : ℝ≥0∞) : + ∫⁻ a, s.indicator (λ _, c) a ∂μ = c * μ s := +by rw [lintegral_indicator _ hs, set_lintegral_const] + lemma set_lintegral_eq_const {f : α → ℝ≥0∞} (hf : measurable f) (r : ℝ≥0∞) : ∫⁻ x in {x | f x = r}, f x ∂μ = r * μ {x | f x = r} := begin @@ -2160,6 +2164,12 @@ lemma set_lintegral_map [measurable_space β] {f : β → ℝ≥0∞} {g : α ∫⁻ y in s, f y ∂(map g μ) = ∫⁻ x in g ⁻¹' s, f (g x) ∂μ := by rw [restrict_map hg hs, lintegral_map hf hg] +lemma lintegral_indicator_const_comp {mβ : measurable_space β} + {f : α → β} {s : set β} (hf : measurable f) (hs : measurable_set s) (c : ℝ≥0∞) : + ∫⁻ a, s.indicator (λ _, c) (f a) ∂μ = c * μ (f ⁻¹' s) := +by rw [lintegral_comp (measurable_const.indicator hs) hf, lintegral_indicator_const hs, + measure.map_apply hf hs] + /-- If `g : α → β` is a measurable embedding and `f : β → ℝ≥0∞` is any function (not necessarily measurable), then `∫⁻ a, f a ∂(map g μ) = ∫⁻ a, f (g a) ∂μ`. Compare with `lintegral_map` wich applies to any measurable `g : α → β` but requires that `f` is measurable as well. -/ diff --git a/src/measure_theory/measure/giry_monad.lean b/src/measure_theory/measure/giry_monad.lean index 44e6288561f38..2439ee3104357 100644 --- a/src/measure_theory/measure/giry_monad.lean +++ b/src/measure_theory/measure/giry_monad.lean @@ -53,6 +53,15 @@ lemma measurable_of_measurable_coe (f : β → measure α) measurable.of_le_map $ supr₂_le $ assume s hs, measurable_space.comap_le_iff_le_map.2 $ by rw [measurable_space.map_comp]; exact h s hs +instance {α : Type*} {m : measurable_space α} : has_measurable_add₂ (measure α) := +begin + refine ⟨measure.measurable_of_measurable_coe _ (λ s hs, _)⟩, + simp_rw [measure.coe_add, pi.add_apply], + refine measurable.add _ _, + { exact (measure.measurable_coe hs).comp measurable_fst, }, + { exact (measure.measurable_coe hs).comp measurable_snd, }, +end + lemma measurable_measure {μ : α → measure β} : measurable μ ↔ ∀ (s : set β) (hs : measurable_set s), measurable (λ b, μ b s) := ⟨λ hμ s hs, (measurable_coe hs).comp hμ, measurable_of_measurable_coe μ⟩ diff --git a/src/measure_theory/measure/measure_space.lean b/src/measure_theory/measure/measure_space.lean index 99c74342f3dc9..fdcf54b5b8a7f 100644 --- a/src/measure_theory/measure/measure_space.lean +++ b/src/measure_theory/measure/measure_space.lean @@ -1784,6 +1784,10 @@ lemma sum_apply_eq_zero' {μ : ι → measure α} {s : set α} (hs : measurable_ sum μ s = 0 ↔ ∀ i, μ i s = 0 := by simp [hs] +lemma sum_comm {ι' : Type*} (μ : ι → ι' → measure α) : + sum (λ n, sum (μ n)) = sum (λ m, sum (λ n, μ n m)) := +by { ext1 s hs, simp_rw [sum_apply _ hs], rw ennreal.tsum_comm, } + lemma ae_sum_iff [countable ι] {μ : ι → measure α} {p : α → Prop} : (∀ᵐ x ∂(sum μ), p x) ↔ ∀ i, ∀ᵐ x ∂(μ i), p x := sum_apply_eq_zero diff --git a/src/probability/kernel/basic.lean b/src/probability/kernel/basic.lean new file mode 100644 index 0000000000000..37c9c54684dc2 --- /dev/null +++ b/src/probability/kernel/basic.lean @@ -0,0 +1,294 @@ +/- +Copyright (c) 2022 Rémy Degenne. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Rémy Degenne +-/ + +import measure_theory.constructions.prod + +/-! +# Markov Kernels + +A kernel from a measurable space `α` to another measurable space `β` is a measurable map +`α → measure β`, where the measurable space instance on `measure β` is the one defined in +`measure_theory.measure.measurable_space`. That is, a kernel `κ` verifies that for all measurable +sets `s` of `β`, `λ a, κ a s` is measurable. + +## Main definitions + +Classes of kernels: +* `kernel α β`: kernels from `α` to `β`, defined as the `add_submonoid` of the measurable + functions in `α → measure β`. +* `is_markov_kernel κ`: a kernel from `α` to `β` is said to be a Markov kernel if for all `a : α`, + `k a` is a probability measure. +* `is_finite_kernel κ`: a kernel from `α` to `β` is said to be finite if there exists `C : ℝ≥0∞` + such that `C < ∞` and for all `a : α`, `κ a univ ≤ C`. This implies in particular that all + measures in the image of `κ` are finite, but is stronger since it requires an uniform bound. This + stronger condition is necessary to ensure that the composition of two finite kernels is finite. +* `is_s_finite_kernel κ`: a kernel is called s-finite if it is a countable sum of finite kernels. + +## Main statements + +* `ext_fun`: if `∫⁻ b, f b ∂(κ a) = ∫⁻ b, f b ∂(η a)` for all measurable functions `f` and all `a`, + then the two kernels `κ` and `η` are equal. + +-/ + +open measure_theory + +open_locale measure_theory ennreal big_operators + +namespace probability_theory + +/-- A kernel from a measurable space `α` to another measurable space `β` is a measurable function +`κ : α → measure β`. The measurable space structure on `measure β` is given by +`measure_theory.measure.measurable_space`. A map `κ : α → measure β` is measurable iff +`∀ s : set β, measurable_set s → measurable (λ a, κ a s)`. -/ +def kernel (α β : Type*) [measurable_space α] [measurable_space β] : + add_submonoid (α → measure β) := +{ carrier := measurable, + zero_mem' := measurable_zero, + add_mem' := λ f g hf hg, measurable.add hf hg, } + +instance {α β : Type*} [measurable_space α] [measurable_space β] : + has_coe_to_fun (kernel α β) (λ _, α → measure β) := ⟨λ κ, κ.val⟩ + +variables {α β ι : Type*} {mα : measurable_space α} {mβ : measurable_space β} + +include mα mβ + +namespace kernel + +@[simp] lemma coe_fn_zero : ⇑(0 : kernel α β) = 0 := rfl +@[simp] lemma coe_fn_add (κ η : kernel α β) : ⇑(κ + η) = κ + η := rfl + +omit mα mβ + +/-- Coercion to a function as an additive monoid homomorphism. -/ +def coe_add_hom (α β : Type*) [measurable_space α] [measurable_space β] : + kernel α β →+ (α → measure β) := +⟨coe_fn, coe_fn_zero, coe_fn_add⟩ + +include mα mβ + +@[simp] lemma zero_apply (a : α) : (0 : kernel α β) a = 0 := rfl + +@[simp] lemma coe_finset_sum (I : finset ι) (κ : ι → kernel α β) : + ⇑(∑ i in I, κ i) = ∑ i in I, κ i := +(coe_add_hom α β).map_sum _ _ + +lemma finset_sum_apply (I : finset ι) (κ : ι → kernel α β) (a : α) : + (∑ i in I, κ i) a = ∑ i in I, κ i a := +by rw [coe_finset_sum, finset.sum_apply] + +lemma finset_sum_apply' (I : finset ι) (κ : ι → kernel α β) (a : α) (s : set β) : + (∑ i in I, κ i) a s = ∑ i in I, κ i a s := +by rw [finset_sum_apply, measure.finset_sum_apply] + +end kernel + +/-- A kernel is a Markov kernel if every measure in its image is a probability measure. -/ +class is_markov_kernel (κ : kernel α β) : Prop := +(is_probability_measure : ∀ a, is_probability_measure (κ a)) + +/-- A kernel is finite if every measure in its image is finite, with a uniform bound. -/ +class is_finite_kernel (κ : kernel α β) : Prop := +(exists_univ_le : ∃ C : ℝ≥0∞, C < ∞ ∧ ∀ a, κ a set.univ ≤ C) + +/-- A constant `C : ℝ≥0∞` such that `C < ∞` (`is_finite_kernel.bound_lt_top κ`) and for all +`a : α` and `s : set β`, `κ a s ≤ C` (`measure_le_bound κ a s`). -/ +noncomputable +def is_finite_kernel.bound (κ : kernel α β) [h : is_finite_kernel κ] : ℝ≥0∞ := +h.exists_univ_le.some + +lemma is_finite_kernel.bound_lt_top (κ : kernel α β) [h : is_finite_kernel κ] : + is_finite_kernel.bound κ < ∞ := +h.exists_univ_le.some_spec.1 + +lemma is_finite_kernel.bound_ne_top (κ : kernel α β) [h : is_finite_kernel κ] : + is_finite_kernel.bound κ ≠ ∞ := +(is_finite_kernel.bound_lt_top κ).ne + +lemma kernel.measure_le_bound (κ : kernel α β) [h : is_finite_kernel κ] (a : α) (s : set β) : + κ a s ≤ is_finite_kernel.bound κ := +(measure_mono (set.subset_univ s)).trans (h.exists_univ_le.some_spec.2 a) + +instance is_finite_kernel_zero (α β : Type*) {mα : measurable_space α} {mβ : measurable_space β} : + is_finite_kernel (0 : kernel α β) := +⟨⟨0, ennreal.coe_lt_top, + λ a, by simp only [kernel.zero_apply, measure.coe_zero, pi.zero_apply, le_zero_iff]⟩⟩ + +instance is_finite_kernel.add (κ η : kernel α β) [is_finite_kernel κ] [is_finite_kernel η] : + is_finite_kernel (κ + η) := +begin + refine ⟨⟨is_finite_kernel.bound κ + is_finite_kernel.bound η, + ennreal.add_lt_top.mpr ⟨is_finite_kernel.bound_lt_top κ, is_finite_kernel.bound_lt_top η⟩, + λ a, _⟩⟩, + simp_rw [kernel.coe_fn_add, pi.add_apply, measure.coe_add, pi.add_apply], + exact add_le_add (kernel.measure_le_bound _ _ _) (kernel.measure_le_bound _ _ _), +end + +variables {κ : kernel α β} + +instance is_markov_kernel.is_probability_measure' [h : is_markov_kernel κ] (a : α) : + is_probability_measure (κ a) := +is_markov_kernel.is_probability_measure a + +instance is_finite_kernel.is_finite_measure [h : is_finite_kernel κ] (a : α) : + is_finite_measure (κ a) := +⟨(kernel.measure_le_bound κ a set.univ).trans_lt (is_finite_kernel.bound_lt_top κ)⟩ + +@[priority 100] +instance is_markov_kernel.is_finite_kernel [h : is_markov_kernel κ] : is_finite_kernel κ := +⟨⟨1, ennreal.one_lt_top, λ a, prob_le_one⟩⟩ + +namespace kernel + +@[ext] lemma ext {κ : kernel α β} {η : kernel α β} (h : ∀ a, κ a = η a) : κ = η := +by { ext1, ext1 a, exact h a, } + +lemma ext_fun {κ η : kernel α β} (h : ∀ a f, measurable f → ∫⁻ b, f b ∂(κ a) = ∫⁻ b, f b ∂(η a)) : + κ = η := +begin + ext a s hs, + specialize h a (s.indicator (λ _, 1)) (measurable.indicator measurable_const hs), + simp_rw [lintegral_indicator_const hs, one_mul] at h, + rw h, +end + +protected lemma measurable (κ : kernel α β) : measurable κ := κ.prop + +protected lemma measurable_coe (κ : kernel α β) {s : set β} (hs : measurable_set s) : + measurable (λ a, κ a s) := +(measure.measurable_coe hs).comp (kernel.measurable κ) + +section sum + +/-- Sum of an indexed family of kernels. -/ +protected noncomputable +def sum [countable ι] (κ : ι → kernel α β) : kernel α β := +{ val := λ a, measure.sum (λ n, κ n a), + property := + begin + refine measure.measurable_of_measurable_coe _ (λ s hs, _), + simp_rw measure.sum_apply _ hs, + exact measurable.ennreal_tsum (λ n, kernel.measurable_coe (κ n) hs), + end, } + +lemma sum_apply [countable ι] (κ : ι → kernel α β) (a : α) : + kernel.sum κ a = measure.sum (λ n, κ n a) := rfl + +lemma sum_apply' [countable ι] (κ : ι → kernel α β) (a : α) {s : set β} (hs : measurable_set s) : + kernel.sum κ a s = ∑' n, κ n a s := +by rw [sum_apply κ a, measure.sum_apply _ hs] + +lemma sum_comm [countable ι] (κ : ι → ι → kernel α β) : + kernel.sum (λ n, kernel.sum (κ n)) = kernel.sum (λ m, kernel.sum (λ n, κ n m)) := +by { ext a s hs, simp_rw [sum_apply], rw measure.sum_comm, } + +@[simp] lemma sum_fintype [fintype ι] (κ : ι → kernel α β) : kernel.sum κ = ∑ i, κ i := +by { ext a s hs, simp only [sum_apply' κ a hs, finset_sum_apply' _ κ a s, tsum_fintype], } + +lemma sum_add [countable ι] (κ η : ι → kernel α β) : + kernel.sum (λ n, κ n + η n) = kernel.sum κ + kernel.sum η := +begin + ext a s hs, + simp only [coe_fn_add, pi.add_apply, sum_apply, measure.sum_apply _ hs, pi.add_apply, + measure.coe_add, tsum_add ennreal.summable ennreal.summable], +end + +end sum + +section s_finite + +/-- A kernel is s-finite if it can be written as the sum of countably many finite kernels. -/ +class is_s_finite_kernel (κ : kernel α β) : Prop := +(tsum_finite : ∃ κs : ℕ → kernel α β, (∀ n, is_finite_kernel (κs n)) ∧ κ = kernel.sum κs) + +@[priority 100] +instance is_finite_kernel.is_s_finite_kernel [h : is_finite_kernel κ] : is_s_finite_kernel κ := +⟨⟨λ n, if n = 0 then κ else 0, + λ n, by { split_ifs, exact h, apply_instance, }, + begin + ext a s hs, + rw kernel.sum_apply' _ _ hs, + have : (λ i, ((ite (i = 0) κ 0) a) s) = λ i, ite (i = 0) (κ a s) 0, + { ext1 i, split_ifs; refl, }, + rw [this, tsum_ite_eq], + end⟩⟩ + +/-- A sequence of finite kernels such that `κ = kernel.sum (seq κ)`. See `is_finite_kernel_seq` +and `kernel_sum_seq`. -/ +noncomputable +def seq (κ : kernel α β) [h : is_s_finite_kernel κ] : + ℕ → kernel α β := +h.tsum_finite.some + +lemma kernel_sum_seq (κ : kernel α β) [h : is_s_finite_kernel κ] : + kernel.sum (seq κ) = κ := +h.tsum_finite.some_spec.2.symm + +lemma measure_sum_seq (κ : kernel α β) [h : is_s_finite_kernel κ] (a : α) : + measure.sum (λ n, seq κ n a) = κ a := +by rw [← kernel.sum_apply, kernel_sum_seq κ] + +instance is_finite_kernel_seq (κ : kernel α β) [h : is_s_finite_kernel κ] (n : ℕ) : + is_finite_kernel (kernel.seq κ n) := +h.tsum_finite.some_spec.1 n + +instance is_s_finite_kernel.add (κ η : kernel α β) [is_s_finite_kernel κ] [is_s_finite_kernel η] : + is_s_finite_kernel (κ + η) := +begin + refine ⟨⟨λ n, seq κ n + seq η n, λ n, infer_instance, _⟩⟩, + rw [sum_add, kernel_sum_seq κ, kernel_sum_seq η], +end + +lemma is_s_finite_kernel.finset_sum {κs : ι → kernel α β} (I : finset ι) + (h : ∀ i ∈ I, is_s_finite_kernel (κs i)) : + is_s_finite_kernel (∑ i in I, κs i) := +begin + classical, + unfreezingI + { induction I using finset.induction with i I hi_nmem_I h_ind h, + { rw [finset.sum_empty], apply_instance, }, + { rw finset.sum_insert hi_nmem_I, + haveI : is_s_finite_kernel (κs i) := h i (finset.mem_insert_self _ _), + haveI : is_s_finite_kernel (∑ (x : ι) in I, κs x), + from h_ind (λ i hiI, h i (finset.mem_insert_of_mem hiI)), + exact is_s_finite_kernel.add _ _, }, }, +end + +lemma is_s_finite_kernel_sum_of_denumerable [denumerable ι] {κs : ι → kernel α β} + (hκs : ∀ n, is_s_finite_kernel (κs n)) : + is_s_finite_kernel (kernel.sum κs) := +begin + let e : ℕ ≃ (ι × ℕ) := denumerable.equiv₂ ℕ (ι × ℕ), + refine ⟨⟨λ n, seq (κs (e n).1) (e n).2, infer_instance, _⟩⟩, + have hκ_eq : kernel.sum κs = kernel.sum (λ n, kernel.sum (seq (κs n))), + { simp_rw kernel_sum_seq, }, + ext a s hs : 2, + rw hκ_eq, + simp_rw kernel.sum_apply' _ _ hs, + change ∑' i m, seq (κs i) m a s = ∑' n, (λ im : ι × ℕ, seq (κs im.fst) im.snd a s) (e n), + rw e.tsum_eq, + { rw tsum_prod' ennreal.summable (λ _, ennreal.summable), }, + { apply_instance, }, +end + +lemma is_s_finite_kernel_sum [countable ι] {κs : ι → kernel α β} + (hκs : ∀ n, is_s_finite_kernel (κs n)) : + is_s_finite_kernel (kernel.sum κs) := +begin + casesI fintype_or_infinite ι, + { rw sum_fintype, + exact is_s_finite_kernel.finset_sum finset.univ (λ i _, hκs i), }, + haveI : encodable ι := encodable.of_countable ι, + haveI : denumerable ι := denumerable.of_encodable_of_infinite ι, + exact is_s_finite_kernel_sum_of_denumerable hκs, +end + +end s_finite + +end kernel + +end probability_theory From 9116dd6709f303dcf781632e15fdef382b0fc579 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Wed, 21 Dec 2022 18:45:58 +0000 Subject: [PATCH 0091/1291] chore(data/set/pointwise): split file and reduce imports (#17991) --- src/algebra/bounds.lean | 1 + src/data/finset/pointwise.lean | 1 + src/data/set/pointwise/basic.lean | 31 +++------------- src/data/set/pointwise/list_of_fn.lean | 49 ++++++++++++++++++++++++++ src/data/set/pointwise/smul.lean | 1 + 5 files changed, 56 insertions(+), 27 deletions(-) create mode 100644 src/data/set/pointwise/list_of_fn.lean diff --git a/src/algebra/bounds.lean b/src/algebra/bounds.lean index dad2177fa3681..9e672062d77e7 100644 --- a/src/algebra/bounds.lean +++ b/src/algebra/bounds.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury G. Kudryashov -/ import algebra.order.group.order_iso +import algebra.order.monoid.order_dual import data.set.pointwise.basic import order.bounds.order_iso import order.conditionally_complete_lattice.basic diff --git a/src/data/finset/pointwise.lean b/src/data/finset/pointwise.lean index 39742d7c1d47d..2863687979276 100644 --- a/src/data/finset/pointwise.lean +++ b/src/data/finset/pointwise.lean @@ -6,6 +6,7 @@ Authors: Floris van Doorn, Yaël Dillies import data.finset.n_ary import data.finset.preimage import data.set.pointwise.smul +import data.set.pointwise.list_of_fn /-! # Pointwise operations of finsets diff --git a/src/data/set/pointwise/basic.lean b/src/data/set/pointwise/basic.lean index a7391f94aeb54..df2bae5a87c48 100644 --- a/src/data/set/pointwise/basic.lean +++ b/src/data/set/pointwise/basic.lean @@ -3,8 +3,11 @@ Copyright (c) 2019 Johan Commelin. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin, Floris van Doorn -/ +import algebra.group_power.basic +import algebra.hom.equiv.basic +import algebra.hom.units import data.set.lattice -import data.list.of_fn +import data.nat.order.basic /-! # Pointwise operations of sets @@ -443,32 +446,6 @@ begin exact ih.trans (subset_mul_right _ hs) } end -@[to_additive] lemma mem_prod_list_of_fn {a : α} {s : fin n → set α} : - a ∈ (list.of_fn s).prod ↔ ∃ f : (Π i : fin n, s i), (list.of_fn (λ i, (f i : α))).prod = a := -begin - induction n with n ih generalizing a, - { simp_rw [list.of_fn_zero, list.prod_nil, fin.exists_fin_zero_pi, eq_comm, set.mem_one] }, - { simp_rw [list.of_fn_succ, list.prod_cons, fin.exists_fin_succ_pi, fin.cons_zero, fin.cons_succ, - mem_mul, @ih, exists_and_distrib_left, exists_exists_eq_and, set_coe.exists, subtype.coe_mk, - exists_prop] } -end - -@[to_additive] lemma mem_list_prod {l : list (set α)} {a : α} : - a ∈ l.prod ↔ ∃ l' : list (Σ s : set α, ↥s), - list.prod (l'.map (λ x, (sigma.snd x : α))) = a ∧ l'.map sigma.fst = l := -begin - induction l using list.of_fn_rec with n f, - simp_rw [list.exists_iff_exists_tuple, list.map_of_fn, list.of_fn_inj', and.left_comm, - exists_and_distrib_left, exists_eq_left, heq_iff_eq, function.comp, mem_prod_list_of_fn], - split, - { rintros ⟨fi, rfl⟩, exact ⟨λ i, ⟨_, fi i⟩, rfl, rfl⟩, }, - { rintros ⟨fi, rfl, rfl⟩, exact ⟨λ i, _, rfl⟩, }, -end - -@[to_additive] lemma mem_pow {a : α} {n : ℕ} : - a ∈ s ^ n ↔ ∃ f : fin n → s, (list.of_fn (λ i, (f i : α))).prod = a := -by rw [←mem_prod_list_of_fn, list.of_fn_const, list.prod_repeat] - @[simp, to_additive] lemma empty_pow {n : ℕ} (hn : n ≠ 0) : (∅ : set α) ^ n = ∅ := by rw [← tsub_add_cancel_of_le (nat.succ_le_of_lt $ nat.pos_of_ne_zero hn), pow_succ, empty_mul] diff --git a/src/data/set/pointwise/list_of_fn.lean b/src/data/set/pointwise/list_of_fn.lean new file mode 100644 index 0000000000000..3356bf4453b14 --- /dev/null +++ b/src/data/set/pointwise/list_of_fn.lean @@ -0,0 +1,49 @@ +/- +Copyright (c) 2022 Eric Wieser. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Eric Wieser +-/ + +import data.set.pointwise.basic +import data.list.of_fn + +/-! +# Pointwise operations with lists of sets + +This file proves some lemmas about pointwise algebraic operations with lists of sets. +-/ + +namespace set + +variables {F α β γ : Type*} +variables [monoid α] {s t : set α} {a : α} {m n : ℕ} + +open_locale pointwise + +@[to_additive] lemma mem_prod_list_of_fn {a : α} {s : fin n → set α} : + a ∈ (list.of_fn s).prod ↔ ∃ f : (Π i : fin n, s i), (list.of_fn (λ i, (f i : α))).prod = a := +begin + induction n with n ih generalizing a, + { simp_rw [list.of_fn_zero, list.prod_nil, fin.exists_fin_zero_pi, eq_comm, set.mem_one] }, + { simp_rw [list.of_fn_succ, list.prod_cons, fin.exists_fin_succ_pi, fin.cons_zero, fin.cons_succ, + mem_mul, @ih, exists_and_distrib_left, exists_exists_eq_and, set_coe.exists, subtype.coe_mk, + exists_prop] } +end + +@[to_additive] lemma mem_list_prod {l : list (set α)} {a : α} : + a ∈ l.prod ↔ ∃ l' : list (Σ s : set α, ↥s), + list.prod (l'.map (λ x, (sigma.snd x : α))) = a ∧ l'.map sigma.fst = l := +begin + induction l using list.of_fn_rec with n f, + simp_rw [list.exists_iff_exists_tuple, list.map_of_fn, list.of_fn_inj', and.left_comm, + exists_and_distrib_left, exists_eq_left, heq_iff_eq, function.comp, mem_prod_list_of_fn], + split, + { rintros ⟨fi, rfl⟩, exact ⟨λ i, ⟨_, fi i⟩, rfl, rfl⟩, }, + { rintros ⟨fi, rfl, rfl⟩, exact ⟨λ i, _, rfl⟩, }, +end + +@[to_additive] lemma mem_pow {a : α} {n : ℕ} : + a ∈ s ^ n ↔ ∃ f : fin n → s, (list.of_fn (λ i, (f i : α))).prod = a := +by rw [←mem_prod_list_of_fn, list.of_fn_const, list.prod_repeat] + +end set diff --git a/src/data/set/pointwise/smul.lean b/src/data/set/pointwise/smul.lean index dd9a12dad6c23..c1e5817916832 100644 --- a/src/data/set/pointwise/smul.lean +++ b/src/data/set/pointwise/smul.lean @@ -6,6 +6,7 @@ Authors: Johan Commelin, Floris van Doorn import algebra.module.basic import data.set.pairwise import data.set.pointwise.basic +import tactic.by_contra /-! # Pointwise operations of sets From a47cda9662ff3925c6df271090b5808adbca5b46 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Wed, 21 Dec 2022 21:36:15 +0000 Subject: [PATCH 0092/1291] refactor(algebra/group_with_zero/defs): use `is_*cancel_mul_zero` (#17963) --- archive/imo/imo1988_q6.lean | 4 +- src/algebra/associated.lean | 28 ++-- src/algebra/char_p/basic.lean | 4 +- src/algebra/group_with_zero/defs.lean | 131 +++++++++--------- src/algebra/ring/regular.lean | 10 +- src/data/nat/basic.lean | 7 +- src/data/nat/choose/basic.lean | 5 +- src/data/nat/gcd/basic.lean | 4 +- src/data/rat/defs.lean | 3 +- .../cyclotomic/primitive_roots.lean | 4 +- src/number_theory/zsqrtd/basic.lean | 5 +- .../localization/fraction_ring.lean | 2 +- .../polynomial/cyclotomic/basic.lean | 2 +- src/ring_theory/valuation/valuation_ring.lean | 3 +- 14 files changed, 92 insertions(+), 120 deletions(-) diff --git a/archive/imo/imo1988_q6.lean b/archive/imo/imo1988_q6.lean index 9260dd93be09a..5db7f6694d4fb 100644 --- a/archive/imo/imo1988_q6.lean +++ b/archive/imo/imo1988_q6.lean @@ -292,6 +292,6 @@ begin { simp only [mul_one, one_mul, add_comm, zero_add] at h, have y_dvd : y ∣ y * k := dvd_mul_right y k, rw [← h, ← add_assoc, nat.dvd_add_left (dvd_mul_left y y)] at y_dvd, - obtain rfl|rfl := (nat.dvd_prime nat.prime_two).mp y_dvd; apply nat.eq_of_mul_eq_mul_left, - exacts [zero_lt_one, h.symm, zero_lt_two, h.symm] } } + obtain rfl|rfl := (nat.dvd_prime nat.prime_two).mp y_dvd; apply mul_left_cancel₀, + exacts [one_ne_zero, h.symm, two_ne_zero, h.symm] } } end diff --git a/src/algebra/associated.lean b/src/algebra/associated.lean index 66ce92bbf4598..b955e8b4be8e8 100644 --- a/src/algebra/associated.lean +++ b/src/algebra/associated.lean @@ -889,22 +889,19 @@ instance : no_zero_divisors (associates α) := have a = 0 ∨ b = 0, from mul_eq_zero.1 this, this.imp (assume h, h.symm ▸ rfl) (assume h, h.symm ▸ rfl))⟩ -lemma eq_of_mul_eq_mul_left : - ∀(a b c : associates α), a ≠ 0 → a * b = a * c → b = c := -begin - rintros ⟨a⟩ ⟨b⟩ ⟨c⟩ ha h, - rcases quotient.exact' h with ⟨u, hu⟩, - have hu : a * (b * ↑u) = a * c, { rwa [← mul_assoc] }, - exact quotient.sound' ⟨u, mul_left_cancel₀ (mk_ne_zero.1 ha) hu⟩ -end - -lemma eq_of_mul_eq_mul_right : - ∀(a b c : associates α), b ≠ 0 → a * b = c * b → a = c := -λ a b c bne0, (mul_comm b a) ▸ (mul_comm b c) ▸ (eq_of_mul_eq_mul_left b a c bne0) +instance : cancel_comm_monoid_with_zero (associates α) := +{ mul_left_cancel_of_ne_zero := + begin + rintros ⟨a⟩ ⟨b⟩ ⟨c⟩ ha h, + rcases quotient.exact' h with ⟨u, hu⟩, + rw [mul_assoc] at hu, + exact quotient.sound' ⟨u, mul_left_cancel₀ (mk_ne_zero.1 ha) hu⟩ + end, + .. (infer_instance : comm_monoid_with_zero (associates α)) } lemma le_of_mul_le_mul_left (a b c : associates α) (ha : a ≠ 0) : a * b ≤ a * c → b ≤ c -| ⟨d, hd⟩ := ⟨d, eq_of_mul_eq_mul_left a _ _ ha $ by rwa ← mul_assoc⟩ +| ⟨d, hd⟩ := ⟨d, mul_left_cancel₀ ha $ by rwa ← mul_assoc⟩ lemma one_or_eq_of_le_of_prime : ∀(p m : associates α), prime p → m ≤ p → (m = 1 ∨ m = p) @@ -922,11 +919,6 @@ match h m d dvd_rfl with or.inl $ bot_unique $ associates.le_of_mul_le_mul_left d m 1 ‹d ≠ 0› this end -instance : cancel_comm_monoid_with_zero (associates α) := -{ mul_left_cancel_of_ne_zero := eq_of_mul_eq_mul_left, - mul_right_cancel_of_ne_zero := eq_of_mul_eq_mul_right, - .. (infer_instance : comm_monoid_with_zero (associates α)) } - instance : canonically_ordered_monoid (associates α) := { exists_mul_of_le := λ a b, id, le_self_mul := λ a b, ⟨b, rfl⟩, diff --git a/src/algebra/char_p/basic.lean b/src/algebra/char_p/basic.lean index f85f9cde38e79..de1806499b205 100644 --- a/src/algebra/char_p/basic.lean +++ b/src/algebra/char_p/basic.lean @@ -411,9 +411,9 @@ or.elim (eq_zero_or_eq_zero_of_mul_eq_zero this) have p ∣ e, from (cast_eq_zero_iff R p e).mp he, have e ∣ p, from dvd_of_mul_left_eq d (eq.symm hmul), have e = p, from dvd_antisymm ‹e ∣ p› ‹p ∣ e›, - have h₀ : p > 0, from gt_of_ge_of_gt hp (nat.zero_lt_succ 1), + have h₀ : 0 < p, from two_pos.trans_le hp, have d * p = 1 * p, by rw ‹e = p› at hmul; rw [one_mul]; exact eq.symm hmul, - show d = 1 ∨ d = p, from or.inl (eq_of_mul_eq_mul_right h₀ this)) + show d = 1 ∨ d = p, from or.inl (mul_right_cancel₀ h₀.ne' this)) section nontrivial diff --git a/src/algebra/group_with_zero/defs.lean b/src/algebra/group_with_zero/defs.lean index c7ad5b9e5d179..7860f5c1884cc 100644 --- a/src/algebra/group_with_zero/defs.lean +++ b/src/algebra/group_with_zero/defs.lean @@ -38,29 +38,75 @@ class mul_zero_class (M₀ : Type*) extends has_mul M₀, has_zero M₀ := (zero_mul : ∀ a : M₀, 0 * a = 0) (mul_zero : ∀ a : M₀, a * 0 = 0) +section mul_zero_class + +variables [mul_zero_class M₀] {a b : M₀} + +@[ematch, simp] lemma zero_mul (a : M₀) : 0 * a = 0 := +mul_zero_class.zero_mul a + +@[ematch, simp] lemma mul_zero (a : M₀) : a * 0 = 0 := +mul_zero_class.mul_zero a + +end mul_zero_class + /-- A mixin for left cancellative multiplication by nonzero elements. -/ @[protect_proj] class is_left_cancel_mul_zero (M₀ : Type u) [has_mul M₀] [has_zero M₀] : Prop := (mul_left_cancel_of_ne_zero : ∀ {a b c : M₀}, a ≠ 0 → a * b = a * c → b = c) +section is_left_cancel_mul_zero + +variables [has_mul M₀] [has_zero M₀] [is_left_cancel_mul_zero M₀] {a b c : M₀} + +lemma mul_left_cancel₀ (ha : a ≠ 0) (h : a * b = a * c) : b = c := +is_left_cancel_mul_zero.mul_left_cancel_of_ne_zero ha h + +lemma mul_right_injective₀ (ha : a ≠ 0) : function.injective ((*) a) := +λ b c, mul_left_cancel₀ ha + +end is_left_cancel_mul_zero + /-- A mixin for right cancellative multiplication by nonzero elements. -/ @[protect_proj] class is_right_cancel_mul_zero (M₀ : Type u) [has_mul M₀] [has_zero M₀] : Prop := (mul_right_cancel_of_ne_zero : ∀ {a b c : M₀}, b ≠ 0 → a * b = c * b → a = c) +section is_right_cancel_mul_zero + +variables [has_mul M₀] [has_zero M₀] [is_right_cancel_mul_zero M₀] {a b c : M₀} + +lemma mul_right_cancel₀ (hb : b ≠ 0) (h : a * b = c * b) : a = c := +is_right_cancel_mul_zero.mul_right_cancel_of_ne_zero hb h + +lemma mul_left_injective₀ (hb : b ≠ 0) : function.injective (λ a, a * b) := +λ a c, mul_right_cancel₀ hb + +end is_right_cancel_mul_zero + /-- A mixin for cancellative multiplication by nonzero elements. -/ @[protect_proj] class is_cancel_mul_zero (M₀ : Type u) [has_mul M₀] [has_zero M₀] extends is_left_cancel_mul_zero M₀, is_right_cancel_mul_zero M₀ : Prop -section mul_zero_class +section comm_semigroup_with_zero -variables [mul_zero_class M₀] {a b : M₀} +variables [comm_semigroup M₀] [has_zero M₀] -@[ematch, simp] lemma zero_mul (a : M₀) : 0 * a = 0 := -mul_zero_class.zero_mul a +lemma is_left_cancel_mul_zero.to_is_right_cancel_mul_zero [is_left_cancel_mul_zero M₀] : + is_right_cancel_mul_zero M₀ := +⟨λ a b c ha h, mul_left_cancel₀ ha $ (mul_comm _ _).trans $ (h.trans (mul_comm _ _))⟩ -@[ematch, simp] lemma mul_zero (a : M₀) : a * 0 = 0 := -mul_zero_class.mul_zero a +lemma is_right_cancel_mul_zero.to_is_left_cancel_mul_zero [is_right_cancel_mul_zero M₀] : + is_left_cancel_mul_zero M₀ := +⟨λ a b c ha h, mul_right_cancel₀ ha $ (mul_comm _ _).trans $ (h.trans (mul_comm _ _))⟩ -end mul_zero_class +lemma is_left_cancel_mul_zero.to_is_cancel_mul_zero [is_left_cancel_mul_zero M₀] : + is_cancel_mul_zero M₀ := +{ .. ‹is_left_cancel_mul_zero M₀›, .. is_left_cancel_mul_zero.to_is_right_cancel_mul_zero } + +lemma is_right_cancel_mul_zero.to_is_cancel_mul_zero [is_right_cancel_mul_zero M₀] : + is_cancel_mul_zero M₀ := +{ .. ‹is_right_cancel_mul_zero M₀›, .. is_right_cancel_mul_zero.to_is_left_cancel_mul_zero } + +end comm_semigroup_with_zero /-- Predicate typeclass for expressing that `a * b = 0` implies `a = 0` or `b = 0` for all `a` and `b` of type `G₀`. -/ @@ -97,31 +143,11 @@ class cancel_monoid_with_zero (M₀ : Type*) extends monoid_with_zero M₀ := (mul_left_cancel_of_ne_zero : ∀ {a b c : M₀}, a ≠ 0 → a * b = a * c → b = c) (mul_right_cancel_of_ne_zero : ∀ {a b c : M₀}, b ≠ 0 → a * b = c * b → a = c) -section cancel_monoid_with_zero - -variables [cancel_monoid_with_zero M₀] {a b c : M₀} - -lemma mul_left_cancel₀ (ha : a ≠ 0) (h : a * b = a * c) : b = c := -cancel_monoid_with_zero.mul_left_cancel_of_ne_zero ha h - -lemma mul_right_cancel₀ (hb : b ≠ 0) (h : a * b = c * b) : a = c := -cancel_monoid_with_zero.mul_right_cancel_of_ne_zero hb h - -lemma mul_right_injective₀ (ha : a ≠ 0) : function.injective ((*) a) := -λ b c, mul_left_cancel₀ ha - -lemma mul_left_injective₀ (hb : b ≠ 0) : function.injective (λ a, a * b) := -λ a c, mul_right_cancel₀ hb - /-- A `cancel_monoid_with_zero` satisfies `is_cancel_mul_zero`. -/ @[priority 100] -instance cancel_monoid_with_zero.to_is_cancel_mul_zero : is_cancel_mul_zero M₀ := -{ mul_left_cancel_of_ne_zero := λ a b c ha h, - cancel_monoid_with_zero.mul_left_cancel_of_ne_zero ha h, - mul_right_cancel_of_ne_zero := λ a b c hb h, - cancel_monoid_with_zero.mul_right_cancel_of_ne_zero hb h, } - -end cancel_monoid_with_zero +instance cancel_monoid_with_zero.to_is_cancel_mul_zero [cancel_monoid_with_zero M₀] : + is_cancel_mul_zero M₀ := +{ .. ‹cancel_monoid_with_zero M₀› } /-- A type `M` is a commutative “monoid with zero” if it is a commutative monoid with zero element, and `0` is left and right absorbing. -/ @@ -132,8 +158,13 @@ class comm_monoid_with_zero (M₀ : Type*) extends comm_monoid M₀, monoid_with `0` is left and right absorbing, and left/right multiplication by a non-zero element is injective. -/ @[protect_proj, ancestor comm_monoid_with_zero cancel_monoid_with_zero] -class cancel_comm_monoid_with_zero (M₀ : Type*) extends - comm_monoid_with_zero M₀, cancel_monoid_with_zero M₀. +class cancel_comm_monoid_with_zero (M₀ : Type*) extends comm_monoid_with_zero M₀ := +(mul_left_cancel_of_ne_zero : ∀ {a b c : M₀}, a ≠ 0 → a * b = a * c → b = c) + +@[priority 100] +instance cancel_comm_monoid_with_zero.to_cancel_monoid_with_zero + [h : cancel_comm_monoid_with_zero M₀] : cancel_monoid_with_zero M₀ := +{ .. h, .. @is_left_cancel_mul_zero.to_is_right_cancel_mul_zero M₀ _ _ { .. h } } /-- A type `G₀` is a “group with zero” if it is a monoid with zero element (distinct from `1`) such that every nonzero element is invertible. @@ -147,42 +178,6 @@ class group_with_zero (G₀ : Type u) extends (inv_zero : (0 : G₀)⁻¹ = 0) (mul_inv_cancel : ∀ a:G₀, a ≠ 0 → a * a⁻¹ = 1) -namespace comm_monoid_with_zero - -variable [comm_monoid_with_zero M₀] - -lemma is_left_cancel_mul_zero.to_is_right_cancel_mul_zero [is_left_cancel_mul_zero M₀] : - is_right_cancel_mul_zero M₀ := -{ mul_right_cancel_of_ne_zero := λ a b c ha h, - begin - rw [mul_comm, mul_comm c] at h, - exact is_left_cancel_mul_zero.mul_left_cancel_of_ne_zero ha h, - end } - -lemma is_right_cancel_mul_zero.to_is_left_cancel_mul_zero [is_right_cancel_mul_zero M₀] : - is_left_cancel_mul_zero M₀ := -{ mul_left_cancel_of_ne_zero := λ a b c ha h, - begin - rw [mul_comm a, mul_comm a c] at h, - exact is_right_cancel_mul_zero.mul_right_cancel_of_ne_zero ha h, - end } - -lemma is_left_cancel_mul_zero.to_is_cancel_mul_zero [is_left_cancel_mul_zero M₀] : - is_cancel_mul_zero M₀ := -{ mul_left_cancel_of_ne_zero := λ _ _ _, - is_left_cancel_mul_zero.mul_left_cancel_of_ne_zero, - mul_right_cancel_of_ne_zero := λ _ _ _, - is_left_cancel_mul_zero.to_is_right_cancel_mul_zero.mul_right_cancel_of_ne_zero } - -lemma is_right_cancel_mul_zero.to_is_cancel_mul_zero [is_right_cancel_mul_zero M₀] : - is_cancel_mul_zero M₀ := -{ mul_left_cancel_of_ne_zero := λ _ _ _, - is_right_cancel_mul_zero.to_is_left_cancel_mul_zero.mul_left_cancel_of_ne_zero, - mul_right_cancel_of_ne_zero := λ _ _ _, - is_right_cancel_mul_zero.mul_right_cancel_of_ne_zero } - -end comm_monoid_with_zero - section group_with_zero variables [group_with_zero G₀] diff --git a/src/algebra/ring/regular.lean b/src/algebra/ring/regular.lean index 81f632d81eb43..7984f9b3f0f13 100644 --- a/src/algebra/ring/regular.lean +++ b/src/algebra/ring/regular.lean @@ -71,18 +71,12 @@ section is_domain @[priority 100] -- see Note [lower instance priority] instance is_domain.to_cancel_monoid_with_zero [semiring α] [is_domain α] : cancel_monoid_with_zero α := -{ mul_left_cancel_of_ne_zero := λ a b c ha h, - is_cancel_mul_zero.mul_left_cancel_of_ne_zero ha h, - mul_right_cancel_of_ne_zero := λ a b c ha h, - is_cancel_mul_zero.mul_right_cancel_of_ne_zero ha h, - .. semiring.to_monoid_with_zero α } +{ .. semiring.to_monoid_with_zero α, .. ‹is_domain α› } variables [comm_semiring α] [is_domain α] @[priority 100] -- see Note [lower instance priority] instance is_domain.to_cancel_comm_monoid_with_zero : cancel_comm_monoid_with_zero α := -{ mul_left_cancel_of_ne_zero := λ a b c ha H, is_domain.mul_left_cancel_of_ne_zero ha H, - mul_right_cancel_of_ne_zero := λ a b c hb H, is_domain.mul_right_cancel_of_ne_zero hb H, - .. (infer_instance : comm_semiring α) } +{ .. ‹comm_semiring α›, .. ‹is_domain α› } end is_domain diff --git a/src/data/nat/basic.lean b/src/data/nat/basic.lean index f4f86b7632b89..4c4b67c7be4c8 100644 --- a/src/data/nat/basic.lean +++ b/src/data/nat/basic.lean @@ -79,15 +79,10 @@ instance : semiring ℕ := infer_instance protected lemma nat.nsmul_eq_mul (m n : ℕ) : m • n = m * n := rfl -theorem nat.eq_of_mul_eq_mul_right {n m k : ℕ} (Hm : 0 < m) (H : n * m = k * m) : n = k := -by rw [mul_comm n m, mul_comm k m] at H; exact nat.eq_of_mul_eq_mul_left Hm H - instance nat.cancel_comm_monoid_with_zero : cancel_comm_monoid_with_zero ℕ := { mul_left_cancel_of_ne_zero := λ _ _ _ h1 h2, nat.eq_of_mul_eq_mul_left (nat.pos_of_ne_zero h1) h2, - mul_right_cancel_of_ne_zero := - λ _ _ _ h1 h2, nat.eq_of_mul_eq_mul_right (nat.pos_of_ne_zero h1) h2, - .. (infer_instance : comm_monoid_with_zero ℕ) } + .. nat.comm_semiring } attribute [simp] nat.not_lt_zero nat.succ_ne_zero nat.succ_ne_self nat.zero_ne_one nat.one_ne_zero diff --git a/src/data/nat/choose/basic.lean b/src/data/nat/choose/basic.lean index 00e1d4f2d4670..fc9d2fcc7e983 100644 --- a/src/data/nat/choose/basic.lean +++ b/src/data/nat/choose/basic.lean @@ -124,9 +124,8 @@ end lemma choose_mul {n k s : ℕ} (hkn : k ≤ n) (hsk : s ≤ k) : n.choose k * k.choose s = n.choose s * (n - s).choose (k - s) := begin - have h : 0 < (n - k)! * (k - s)! * s! := - mul_pos (mul_pos (factorial_pos _) (factorial_pos _)) (factorial_pos _), - refine eq_of_mul_eq_mul_right h _, + have h : (n - k)! * (k - s)! * s! ≠ 0, by apply_rules [mul_ne_zero, factorial_ne_zero], + refine mul_right_cancel₀ h _, calc n.choose k * k.choose s * ((n - k)! * (k - s)! * s!) = n.choose k * (k.choose s * s! * (k - s)!) * (n - k)! diff --git a/src/data/nat/gcd/basic.lean b/src/data/nat/gcd/basic.lean index de73719b823cd..fd5e2cf322c9a 100644 --- a/src/data/nat/gcd/basic.lean +++ b/src/data/nat/gcd/basic.lean @@ -104,9 +104,9 @@ end theorem gcd_div {m n k : ℕ} (H1 : k ∣ m) (H2 : k ∣ n) : gcd (m / k) (n / k) = gcd m n / k := -or.elim (nat.eq_zero_or_pos k) +(decidable.eq_or_ne k 0).elim (λk0, by rw [k0, nat.div_zero, nat.div_zero, nat.div_zero, gcd_zero_right]) - (λH3, nat.eq_of_mul_eq_mul_right H3 $ by rw [ + (λH3, mul_right_cancel₀ H3 $ by rw [ nat.div_mul_cancel (dvd_gcd H1 H2), ←gcd_mul_right, nat.div_mul_cancel H1, nat.div_mul_cancel H2]) diff --git a/src/data/rat/defs.lean b/src/data/rat/defs.lean index b3059d00bf39e..0419d00033201 100644 --- a/src/data/rat/defs.lean +++ b/src/data/rat/defs.lean @@ -162,8 +162,7 @@ begin all_goals { exact int.coe_nat_dvd.2 (nat.gcd_dvd_left _ _) } }, intros a c h, suffices bd : b / a.gcd b = d / c.gcd d, - { refine ⟨_, bd⟩, - apply nat.eq_of_mul_eq_mul_left hb, + { refine ⟨mul_left_cancel₀ hb.ne' _, bd⟩, rw [← nat.mul_div_assoc _ (nat.gcd_dvd_left _ _), mul_comm, nat.mul_div_assoc _ (nat.gcd_dvd_right _ _), bd, ← nat.mul_div_assoc _ (nat.gcd_dvd_right _ _), h, mul_comm, diff --git a/src/number_theory/cyclotomic/primitive_roots.lean b/src/number_theory/cyclotomic/primitive_roots.lean index a6dc27d8567df..a2865d33a9a1e 100644 --- a/src/number_theory/cyclotomic/primitive_roots.lean +++ b/src/number_theory/cyclotomic/primitive_roots.lean @@ -362,12 +362,12 @@ begin pnat.pow_coe, pnat.pow_coe, nat.totient_prime_pow hpri.out (k - s).succ_pos, nat.totient_prime_pow hpri.out k.succ_pos, mul_comm _ (↑p - 1), mul_assoc, mul_comm (↑p ^ (k.succ - 1))] at this, - replace this := nat.eq_of_mul_eq_mul_left (tsub_pos_iff_lt.2 (nat.prime.one_lt hpri.out)) this, + replace this := mul_left_cancel₀ (tsub_pos_iff_lt.2 hpri.out.one_lt).ne' this, have Hex : k.succ - 1 = (k - s).succ - 1 + s, { simp only [nat.succ_sub_succ_eq_sub, tsub_zero], exact (nat.sub_add_cancel hs).symm }, rw [Hex, pow_add] at this, - exact nat.eq_of_mul_eq_mul_left (pow_pos hpri.out.pos _) this }, + exact mul_left_cancel₀ (pow_ne_zero _ hpri.out.ne_zero) this }, all_goals { apply_instance } end diff --git a/src/number_theory/zsqrtd/basic.lean b/src/number_theory/zsqrtd/basic.lean index b86a6d38ba44b..4e29530bb4f37 100644 --- a/src/number_theory/zsqrtd/basic.lean +++ b/src/number_theory/zsqrtd/basic.lean @@ -231,8 +231,7 @@ protected lemma eq_of_smul_eq_smul_left {a : ℤ} {b c : ℤ√d} begin rw ext at h ⊢, apply and.imp _ _ h; - { simp only [smul_re, smul_im], - exact int.eq_of_mul_eq_mul_left ha }, + { simpa only [smul_re, smul_im] using mul_left_cancel₀ ha }, end section gcd @@ -641,7 +640,7 @@ let g := x.gcd y in or.elim g.eq_zero_or_pos let ⟨m, n, co, (hx : x = m * g), (hy : y = n * g)⟩ := nat.exists_coprime gpos in begin rw [hx, hy] at h, - have : m * m = d * (n * n) := nat.eq_of_mul_eq_mul_left (mul_pos gpos gpos) + have : m * m = d * (n * n) := mul_left_cancel₀ (mul_pos gpos gpos).ne' (by simpa [mul_comm, mul_left_comm] using h), have co2 := let co1 := co.mul_right co in co1.mul co1, exact nonsquare.ns d m (nat.dvd_antisymm (by rw this; apply dvd_mul_right) $ diff --git a/src/ring_theory/localization/fraction_ring.lean b/src/ring_theory/localization/fraction_ring.lean index 44078b93fb171..f9b263c6817a5 100644 --- a/src/ring_theory/localization/fraction_ring.lean +++ b/src/ring_theory/localization/fraction_ring.lean @@ -57,7 +57,7 @@ instance rat.is_fraction_ring : is_fraction_ring ℤ ℚ := rw [eq_int_cast, eq_int_cast, int.cast_inj], refine ⟨by { rintro rfl, use 1 }, _⟩, rintro ⟨⟨c, hc⟩, h⟩, - apply int.eq_of_mul_eq_mul_right _ h, + apply mul_right_cancel₀ _ h, rwa mem_non_zero_divisors_iff_ne_zero at hc, end } diff --git a/src/ring_theory/polynomial/cyclotomic/basic.lean b/src/ring_theory/polynomial/cyclotomic/basic.lean index 836c87d3708fa..7ec13bdc53eed 100644 --- a/src/ring_theory/polynomial/cyclotomic/basic.lean +++ b/src/ring_theory/polynomial/cyclotomic/basic.lean @@ -881,7 +881,7 @@ begin rw [polynomial.map_mul, map_cyclotomic_int, map_cyclotomic_int, map_expand, map_cyclotomic_int], refine is_coprime.mul_dvd (cyclotomic.is_coprime_rat (λ h, _)) _ _, { replace h : n * p = n * 1 := by simp [h], - exact nat.prime.ne_one hp (nat.eq_of_mul_eq_mul_left hnpos h) }, + exact nat.prime.ne_one hp (mul_left_cancel₀ hnpos.ne' h) }, { have hpos : 0 < n * p := mul_pos hnpos hp.pos, have hprim := complex.is_primitive_root_exp _ hpos.ne', rw [cyclotomic_eq_minpoly_rat hprim hpos], diff --git a/src/ring_theory/valuation/valuation_ring.lean b/src/ring_theory/valuation/valuation_ring.lean index 6631c1583974d..979cc5c29a425 100644 --- a/src/ring_theory/valuation/valuation_ring.lean +++ b/src/ring_theory/valuation/valuation_ring.lean @@ -129,8 +129,7 @@ instance : linear_ordered_comm_group_with_zero (value_group A K) := rw [← mul_smul, algebra.smul_def] at hf, nth_rewrite 1 ← one_mul b at hf, rw ← (algebra_map A K).map_one at hf, - exact is_fraction_ring.injective _ _ - (cancel_comm_monoid_with_zero.mul_right_cancel_of_ne_zero hb hf).symm }, + exact is_fraction_ring.injective _ _ (mul_right_cancel₀ hb hf).symm }, apply quotient.sound', use [this.unit, rfl], end, From 026bed67f28f47a5e2776e3b87cb25cc0396b57b Mon Sep 17 00:00:00 2001 From: Anatole Dedecker Date: Thu, 22 Dec 2022 03:16:16 +0000 Subject: [PATCH 0093/1291] feat(analysis/normed_space/compact_operator): generalize `is_closed_set_of_compact_operator` (#17715) --- .../normed_space/compact_operator.lean | 62 +++++++++---------- src/topology/algebra/group/basic.lean | 23 +++++++ 2 files changed, 52 insertions(+), 33 deletions(-) diff --git a/src/analysis/normed_space/compact_operator.lean b/src/analysis/normed_space/compact_operator.lean index a2e78130ad594..28632b1ffe179 100644 --- a/src/analysis/normed_space/compact_operator.lean +++ b/src/analysis/normed_space/compact_operator.lean @@ -38,11 +38,6 @@ maps. Instead we define it as a predicate over bare functions, although it reall for linear functions, because Lean is really good at finding coercions to bare functions (whereas coercing from continuous linear maps to linear maps often needs type ascriptions). -## TODO - -Once we have the strong operator topology on spaces of linear maps between two TVSs, -`is_closed_set_of_is_compact_operator` should be generalized to this setup. - ## References * Bourbaki, *Spectral Theory*, chapters 3 to 5, to be published (2022) @@ -375,56 +370,57 @@ hf end continuous +/-- The set of compact operators from a normed space to a complete topological vector space is +closed. -/ lemma is_closed_set_of_is_compact_operator {𝕜₁ 𝕜₂ : Type*} [nontrivially_normed_field 𝕜₁] - [nontrivially_normed_field 𝕜₂] {σ₁₂ : 𝕜₁ →+* 𝕜₂} [ring_hom_isometric σ₁₂] {M₁ M₂ : Type*} - [seminormed_add_comm_group M₁] [normed_add_comm_group M₂] [normed_space 𝕜₁ M₁] - [normed_space 𝕜₂ M₂] [complete_space M₂] : + [normed_field 𝕜₂] {σ₁₂ : 𝕜₁ →+* 𝕜₂} {M₁ M₂ : Type*} [seminormed_add_comm_group M₁] + [add_comm_group M₂] [normed_space 𝕜₁ M₁] [module 𝕜₂ M₂] [uniform_space M₂] [uniform_add_group M₂] + [has_continuous_const_smul 𝕜₂ M₂] [t2_space M₂] [complete_space M₂] : is_closed {f : M₁ →SL[σ₁₂] M₂ | is_compact_operator f} := begin refine is_closed_of_closure_subset _, rintros u hu, - rw metric.mem_closure_iff at hu, + rw [mem_closure_iff_nhds_zero] at hu, suffices : totally_bounded (u '' metric.closed_ball 0 1), { change is_compact_operator (u : M₁ →ₛₗ[σ₁₂] M₂), rw is_compact_operator_iff_is_compact_closure_image_closed_ball (u : M₁ →ₛₗ[σ₁₂] M₂) zero_lt_one, exact is_compact_of_totally_bounded_is_closed this.closure is_closed_closure }, - rw metric.totally_bounded_iff, - intros ε hε, - rcases hu (ε/2) (by linarith) with ⟨v, hv, huv⟩, - rcases (hv.is_compact_closure_image_closed_ball 1).finite_cover_balls - (show 0 < ε/2, by linarith) with ⟨T, -, hT, hTv⟩, + rw totally_bounded_iff_subset_finite_Union_nhds_zero, + intros U hU, + rcases exists_nhds_zero_half hU with ⟨V, hV, hVU⟩, + let SV : set M₁ × set M₂ := ⟨closed_ball 0 1, -V⟩, + rcases hu {f | ∀ x ∈ SV.1, f x ∈ SV.2} (continuous_linear_map.has_basis_nhds_zero.mem_of_mem + ⟨normed_space.is_vonN_bounded_closed_ball _ _ _, neg_mem_nhds_zero M₂ hV⟩) with ⟨v, hv, huv⟩, + rcases totally_bounded_iff_subset_finite_Union_nhds_zero.mp + (hv.is_compact_closure_image_closed_ball 1).totally_bounded V hV with ⟨T, hT, hTv⟩, have hTv : v '' closed_ball 0 1 ⊆ _ := subset_closure.trans hTv, refine ⟨T, hT, _⟩, - rw image_subset_iff at ⊢ hTv, + rw [image_subset_iff, preimage_Union₂] at ⊢ hTv, intros x hx, specialize hTv hx, - rw [mem_preimage, mem_Union₂] at ⊢ hTv, + rw [mem_Union₂] at ⊢ hTv, rcases hTv with ⟨t, ht, htx⟩, refine ⟨t, ht, _⟩, - suffices : dist (u x) (v x) < ε/2, - { rw mem_ball at *, - linarith [dist_triangle (u x) (v x) t] }, - rw mem_closed_ball_zero_iff at hx, - calc dist (u x) (v x) - = ‖u x - v x‖ : dist_eq_norm _ _ - ... = ‖(u - v) x‖ : by rw continuous_linear_map.sub_apply; refl - ... ≤ ‖u - v‖ : (u - v).unit_le_op_norm x hx - ... = dist u v : (dist_eq_norm _ _).symm - ... < ε/2 : huv + rw [mem_preimage, mem_vadd_set_iff_neg_vadd_mem, vadd_eq_add, neg_add_eq_sub] at ⊢ htx, + convert hVU _ htx _ (huv x hx) using 1, + rw [continuous_linear_map.sub_apply], + abel end lemma compact_operator_topological_closure {𝕜₁ 𝕜₂ : Type*} [nontrivially_normed_field 𝕜₁] - [nontrivially_normed_field 𝕜₂] {σ₁₂ : 𝕜₁ →+* 𝕜₂} [ring_hom_isometric σ₁₂] {M₁ M₂ : Type*} - [seminormed_add_comm_group M₁] [normed_add_comm_group M₂] [normed_space 𝕜₁ M₁] - [normed_space 𝕜₂ M₂] [complete_space M₂] : + [normed_field 𝕜₂] {σ₁₂ : 𝕜₁ →+* 𝕜₂} {M₁ M₂ : Type*} + [seminormed_add_comm_group M₁] [add_comm_group M₂] [normed_space 𝕜₁ M₁] [module 𝕜₂ M₂] + [uniform_space M₂] [uniform_add_group M₂] [has_continuous_const_smul 𝕜₂ M₂] [t2_space M₂] + [complete_space M₂] [has_continuous_smul 𝕜₂ (M₁ →SL[σ₁₂] M₂)] : (compact_operator σ₁₂ M₁ M₂).topological_closure = compact_operator σ₁₂ M₁ M₂ := set_like.ext' (is_closed_set_of_is_compact_operator.closure_eq) lemma is_compact_operator_of_tendsto {ι 𝕜₁ 𝕜₂ : Type*} [nontrivially_normed_field 𝕜₁] - [nontrivially_normed_field 𝕜₂] {σ₁₂ : 𝕜₁ →+* 𝕜₂} [ring_hom_isometric σ₁₂] {M₁ M₂ : Type*} - [seminormed_add_comm_group M₁] [normed_add_comm_group M₂] [normed_space 𝕜₁ M₁] - [normed_space 𝕜₂ M₂] [complete_space M₂] {l : filter ι} [l.ne_bot] {F : ι → M₁ →SL[σ₁₂] M₂} - {f : M₁ →SL[σ₁₂] M₂} (hf : tendsto F l (𝓝 f)) (hF : ∀ᶠ i in l, is_compact_operator (F i)) : + [normed_field 𝕜₂] {σ₁₂ : 𝕜₁ →+* 𝕜₂} {M₁ M₂ : Type*} + [seminormed_add_comm_group M₁] [add_comm_group M₂] [normed_space 𝕜₁ M₁] [module 𝕜₂ M₂] + [uniform_space M₂] [uniform_add_group M₂] [has_continuous_const_smul 𝕜₂ M₂] [t2_space M₂] + [complete_space M₂] {l : filter ι} [l.ne_bot] {F : ι → M₁ →SL[σ₁₂] M₂} {f : M₁ →SL[σ₁₂] M₂} + (hf : tendsto F l (𝓝 f)) (hF : ∀ᶠ i in l, is_compact_operator (F i)) : is_compact_operator f := is_closed_set_of_is_compact_operator.mem_of_tendsto hf hF diff --git a/src/topology/algebra/group/basic.lean b/src/topology/algebra/group/basic.lean index 9a7b147a6b09b..24df65abf3824 100644 --- a/src/topology/algebra/group/basic.lean +++ b/src/topology/algebra/group/basic.lean @@ -475,6 +475,14 @@ variable (G) lemma nhds_one_symm : comap has_inv.inv (𝓝 (1 : G)) = 𝓝 (1 : G) := ((homeomorph.inv G).comap_nhds_eq _).trans (congr_arg nhds inv_one) +@[to_additive] +lemma nhds_one_symm' : map has_inv.inv (𝓝 (1 : G)) = 𝓝 (1 : G) := +((homeomorph.inv G).map_nhds_eq _).trans (congr_arg nhds inv_one) + +@[to_additive] +lemma inv_mem_nhds_one {S : set G} (hS : S ∈ (𝓝 1 : filter G)) : S⁻¹ ∈ (𝓝 (1 : G)) := +by rwa [← nhds_one_symm'] at hS + /-- The map `(x, y) ↦ (x, xy)` as a homeomorphism. This is a shear mapping. -/ @[to_additive "The map `(x, y) ↦ (x, x + y)` as a homeomorphism. This is a shear mapping."] @@ -615,6 +623,21 @@ lemma nhds_translation_mul_inv (x : G) : comap (λ y : G, y * x⁻¹) (𝓝 1) = @[to_additive] lemma map_mul_left_nhds_one (x : G) : map ((*) x) (𝓝 1) = 𝓝 x := by simp +@[to_additive] lemma filter.has_basis.nhds_of_one {ι : Sort*} {p : ι → Prop} {s : ι → set G} + (hb : has_basis (𝓝 1 : filter G) p s) (x : G) : has_basis (𝓝 x) p (λ i, {y | y / x ∈ s i}) := +begin + rw ← nhds_translation_mul_inv, + simp_rw [div_eq_mul_inv], + exact hb.comap _ +end + +@[to_additive] lemma mem_closure_iff_nhds_one {x : G} {s : set G} : + x ∈ closure s ↔ ∀ U ∈ (𝓝 1 : filter G), ∃ y ∈ s, y / x ∈ U := +begin + rw mem_closure_iff_nhds_basis ((𝓝 1 : filter G).basis_sets.nhds_of_one x), + refl +end + /-- A monoid homomorphism (a bundled morphism of a type that implements `monoid_hom_class`) from a topological group to a topological monoid is continuous provided that it is continuous at one. See also `uniform_continuous_of_continuous_at_one`. -/ From 207cfac9fcd06138865b5d04f7091e46d9320432 Mon Sep 17 00:00:00 2001 From: Felix-Weilacher Date: Thu, 22 Dec 2022 07:41:58 +0000 Subject: [PATCH 0094/1291] feat(measure_theory/measurable_space) add measurable Schroeder-Bernstein (#17988) Add a version of the Schroeder-Bernstein theorem for measurable spaces, embeddings, and equivalences. --- src/measure_theory/measurable_space.lean | 69 +++++++++++++++++++++++- 1 file changed, 67 insertions(+), 2 deletions(-) diff --git a/src/measure_theory/measurable_space.lean b/src/measure_theory/measurable_space.lean index 821af074c9ce5..b240ba194da85 100644 --- a/src/measure_theory/measurable_space.lean +++ b/src/measure_theory/measurable_space.lean @@ -1248,11 +1248,18 @@ def pi_equiv_pi_subtype_prod (p : δ' → Prop) [decidable_pred p] : measurable_to_fun := measurable_pi_equiv_pi_subtype_prod π p, measurable_inv_fun := measurable_pi_equiv_pi_subtype_prod_symm π p } +/-- If `s` is a measurable set in a measurable space, that space is equivalent +to the sum of `s` and `sᶜ`.-/ +def sum_compl {s : set α} [decidable_pred s] (hs : measurable_set s) : s ⊕ (sᶜ : set α) ≃ᵐ α := +{ to_equiv := sum_compl s, + measurable_to_fun := by {apply measurable.sum_elim; exact measurable_subtype_coe}, + measurable_inv_fun := measurable.dite measurable_inl measurable_inr hs } + end measurable_equiv namespace measurable_embedding -variables [measurable_space α] [measurable_space β] [measurable_space γ] {f : α → β} +variables [measurable_space α] [measurable_space β] [measurable_space γ] {f : α → β} {g : β → α} /-- A set is equivalent to its image under a function `f` as measurable spaces, if `f` is a measurable embedding -/ @@ -1283,11 +1290,69 @@ begin exact (measurable_embedding.subtype_coe hf₂).comp e.measurable_embedding end -lemma of_measurable_inverse {g : β → α} (hf₁ : measurable f) +lemma of_measurable_inverse (hf₁ : measurable f) (hf₂ : measurable_set (range f)) (hg : measurable g) (H : left_inverse g f) : measurable_embedding f := of_measurable_inverse_on_range hf₁ hf₂ (hg.comp measurable_subtype_coe) H +open_locale classical + +/-- The **`measurable Schröder-Bernstein Theorem**: Given measurable embeddings +`α → β` and `β → α`, we can find a measurable equivalence `α ≃ᵐ β`.-/ +noncomputable +def schroeder_bernstein {f : α → β} {g : β → α} + (hf : measurable_embedding f)(hg : measurable_embedding g) : α ≃ᵐ β := +begin + let F : set α → set α := λ A, (g '' (f '' A)ᶜ)ᶜ, + -- We follow the proof of the usual SB theorem in mathlib, + -- the crux of which is finding a fixed point of this F. + -- However, we must find this fixed point manually instead of invoking Knaster-Tarski + -- in order to make sure it is measurable. + suffices : Σ' A : set α, measurable_set A ∧ F A = A, + { rcases this with ⟨A, Ameas, Afp⟩, + let B := f '' A, + have Bmeas : measurable_set B := hf.measurable_set_image' Ameas, + refine (measurable_equiv.sum_compl Ameas).symm.trans + (measurable_equiv.trans _ (measurable_equiv.sum_compl Bmeas)), + apply measurable_equiv.sum_congr (hf.equiv_image _), + have : Aᶜ = g '' Bᶜ, + { apply compl_injective, + rw ← Afp, + simp, }, + rw this, + exact (hg.equiv_image _).symm, }, + have Fmono : ∀ {A B}, A ⊆ B → F A ⊆ F B := λ A B hAB, + compl_subset_compl.mpr $ set.image_subset _ $ + compl_subset_compl.mpr $ set.image_subset _ hAB, + let X : ℕ → set α := λ n, F^[n] univ, + refine ⟨Inter X, _, _⟩, + { apply measurable_set.Inter, + intros n, + induction n with n ih, + { exact measurable_set.univ }, + rw [function.iterate_succ', function.comp_apply], + exact (hg.measurable_set_image' (hf.measurable_set_image' ih).compl).compl, }, + apply subset_antisymm, + { apply subset_Inter, + intros n, + cases n, + { exact subset_univ _ }, + rw [function.iterate_succ', function.comp_apply], + exact Fmono (Inter_subset _ _ ), }, + rintros x hx ⟨y, hy, rfl⟩, + rw mem_Inter at hx, + apply hy, + rw (inj_on_of_injective hf.injective _).image_Inter_eq, + swap, { apply_instance }, + rw mem_Inter, + intro n, + specialize hx n.succ, + rw [function.iterate_succ', function.comp_apply] at hx, + by_contradiction h, + apply hx, + exact ⟨y, h, rfl⟩, +end + end measurable_embedding namespace filter From d55a93d80890ba432cb9a5455636653ae1faeedd Mon Sep 17 00:00:00 2001 From: datokrat Date: Thu, 22 Dec 2022 09:08:42 +0000 Subject: [PATCH 0095/1291] feat(linear_algebra/affine_space/affine_subspace): induction principle (#17520) Proves affine analogues of `submodule.span_induction'`, `submodule.span_induction'` and `submodule.span_span_coe_preimage`. --- .../affine_space/affine_subspace.lean | 41 +++++++++++++++++++ 1 file changed, 41 insertions(+) diff --git a/src/linear_algebra/affine_space/affine_subspace.lean b/src/linear_algebra/affine_space/affine_subspace.lean index 51a3e02811e32..83e3fddab70b9 100644 --- a/src/linear_algebra/affine_space/affine_subspace.lean +++ b/src/linear_algebra/affine_space/affine_subspace.lean @@ -1136,6 +1136,47 @@ by rw [←not_iff_not, ←ne.def, ←ne.def, ←nonempty_iff_ne_bot, affine_span variables {k} +/-- +An induction principle for span membership. If `p` holds for all elements of `s` and is +preserved under certain affine combinations, then `p` holds for all elements of the span of `s`. +-/ +lemma affine_span_induction {x : P} {s : set P} {p : P → Prop} (h : x ∈ affine_span k s) + (Hs : ∀ x : P, x ∈ s → p x) + (Hc : ∀ (c : k) (u v w : P), p u → p v → p w → p (c • (u -ᵥ v) +ᵥ w)) : p x := +(@affine_span_le _ _ _ _ _ _ _ _ ⟨p, Hc⟩).mpr Hs h + +/-- A dependent version of `affine_span_induction`. -/ +lemma affine_span_induction' {s : set P} {p : Π x, x ∈ affine_span k s → Prop} + (Hs : ∀ y (hys : y ∈ s), p y (subset_affine_span k _ hys)) + (Hc : ∀ (c : k) u hu v hv w hw, p u hu → p v hv → p w hw → + p (c • (u -ᵥ v) +ᵥ w) (affine_subspace.smul_vsub_vadd_mem _ _ hu hv hw)) + {x : P} (h : x ∈ affine_span k s) : p x h := +begin + refine exists.elim _ (λ (hx : x ∈ affine_span k s) (hc : p x hx), hc), + refine @affine_span_induction k V P _ _ _ _ _ _ _ h _ _, + { exact (λ y hy, ⟨subset_affine_span _ _ hy, Hs y hy⟩) }, + { exact (λ c u v w hu hv hw, exists.elim hu $ λ hu' hu, exists.elim hv $ λ hv' hv, + exists.elim hw $ λ hw' hw, + ⟨affine_subspace.smul_vsub_vadd_mem _ _ hu' hv' hw', Hc _ _ _ _ _ _ _ hu hv hw⟩) }, +end + +section with_local_instance + +local attribute [instance] affine_subspace.to_add_torsor + +/-- A set, considered as a subset of its spanned affine subspace, spans the whole subspace. -/ +@[simp] lemma affine_span_coe_preimage_eq_top (A : set P) [nonempty A] : + affine_span k ((coe : affine_span k A → P) ⁻¹' A) = ⊤ := +begin + rw [eq_top_iff], + rintro ⟨x, hx⟩ -, + refine affine_span_induction' (λ y hy, _) (λ c u hu v hv w hw, _) hx, + { exact subset_affine_span _ _ hy }, + { exact affine_subspace.smul_vsub_vadd_mem _ _ }, +end + +end with_local_instance + /-- Suppose a set of vectors spans `V`. Then a point `p`, together with those vectors added to `p`, spans `P`. -/ lemma affine_span_singleton_union_vadd_eq_top_of_span_eq_top {s : set V} (p : P) From 412b1515f61ae38ad816d57382f96f4bcd226d08 Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Thu, 22 Dec 2022 09:08:43 +0000 Subject: [PATCH 0096/1291] feat(group_theory/order_of_element): a condition on o(x) and o(y) that implies o(xy)=o(y) (#17997) + The main theorem is `order_of_mul_eq_right_of_forall_prime_mul_dvd`, which depends on `order_of_dvd_lcm_mul`, a variant of `order_of_mul_dvd_lcm`, and `dvd_of_forall_prime_mul_dvd`. + Also add lemmas `factorization_prime_le_iff_dvd` and `factorization_lcm` that were used in another approach towards the same theorem. Co-authored-by: Oliver Nash --- src/data/nat/factorization/basic.lean | 27 ++++++++++--- src/data/nat/prime.lean | 8 ++++ src/group_theory/order_of_element.lean | 56 ++++++++++++++++++++++---- 3 files changed, 77 insertions(+), 14 deletions(-) diff --git a/src/data/nat/factorization/basic.lean b/src/data/nat/factorization/basic.lean index 6c4a33df60166..9fa837a0571fe 100644 --- a/src/data/nat/factorization/basic.lean +++ b/src/data/nat/factorization/basic.lean @@ -387,6 +387,14 @@ begin { rintro ⟨c, rfl⟩, rw factorization_mul hd (right_ne_zero_of_mul hn), simp }, end +lemma factorization_prime_le_iff_dvd {d n : ℕ} (hd : d ≠ 0) (hn : n ≠ 0) : + (∀ p : ℕ, p.prime → d.factorization p ≤ n.factorization p) ↔ d ∣ n := +begin + rw ← factorization_le_iff_dvd hd hn, + refine ⟨λ h p, (em p.prime).elim (h p) (λ hp, _), λ h p _, h p⟩, + simp_rw factorization_eq_zero_of_non_prime _ hp, +end + lemma pow_succ_factorization_not_dvd {n p : ℕ} (hn : n ≠ 0) (hp : p.prime) : ¬ p ^ (n.factorization p + 1) ∣ n := begin @@ -558,13 +566,11 @@ begin rcases eq_or_ne n 0 with rfl | hn, { simp }, rcases eq_or_ne d 0 with rfl | hd, { simp only [zero_dvd_iff, hn, false_iff, not_forall], - refine ⟨2, n, prime_two, ⟨dvd_zero _, _⟩⟩, - apply mt (le_of_dvd hn.bot_lt) (not_le.mpr (lt_two_pow n)) }, + exact ⟨2, n, prime_two, dvd_zero _, mt (le_of_dvd hn.bot_lt) (lt_two_pow n).not_le⟩ }, refine ⟨λ h p k _ hpkd, dvd_trans hpkd h, _⟩, - rw [←factorization_le_iff_dvd hd hn, finsupp.le_def], - intros h p, - by_cases pp : prime p, swap, { simp [factorization_eq_zero_of_non_prime d pp] }, - rw ←pp.pow_dvd_iff_le_factorization hn, + rw [←factorization_prime_le_iff_dvd hd hn], + intros h p pp, + simp_rw ←pp.pow_dvd_iff_le_factorization hn, exact h p _ pp (ord_proj_dvd _ _) end @@ -597,6 +603,15 @@ begin simp [←factorization_le_iff_dvd he_pos hd_pos, h1, hea', heb'] }, end +lemma factorization_lcm {a b : ℕ} (ha : a ≠ 0) (hb : b ≠ 0) : + (a.lcm b).factorization = a.factorization ⊔ b.factorization := +begin + rw [← add_right_inj (a.gcd b).factorization, + ← factorization_mul (mt gcd_eq_zero_iff.1 $ λ h, ha h.1) (lcm_ne_zero ha hb), + gcd_mul_lcm, factorization_gcd ha hb, factorization_mul ha hb], + ext1, exact (min_add_max _ _).symm, +end + @[to_additive sum_factors_gcd_add_sum_factors_mul] lemma prod_factors_gcd_mul_prod_factors_mul {β : Type*} [comm_monoid β] (m n : ℕ) (f : ℕ → β) : (m.gcd n).factors.to_finset.prod f * (m * n).factors.to_finset.prod f diff --git a/src/data/nat/prime.lean b/src/data/nat/prime.lean index de9e265131102..261e2de73a035 100644 --- a/src/data/nat/prime.lean +++ b/src/data/nat/prime.lean @@ -402,6 +402,14 @@ theorem exists_dvd_of_not_prime2 {n : ℕ} (n2 : 2 ≤ n) (np : ¬ prime n) : theorem exists_prime_and_dvd {n : ℕ} (hn : n ≠ 1) : ∃ p, prime p ∧ p ∣ n := ⟨min_fac n, min_fac_prime hn, min_fac_dvd _⟩ +theorem dvd_of_forall_prime_mul_dvd {a b : ℕ} + (hdvd : ∀ p : ℕ, p.prime → p ∣ a → p * a ∣ b) : a ∣ b := +begin + obtain rfl | ha := eq_or_ne a 1, { apply one_dvd }, + obtain ⟨p, hp⟩ := exists_prime_and_dvd ha, + exact trans (dvd_mul_left a p) (hdvd p hp.1 hp.2), +end + /-- Euclid's theorem on the **infinitude of primes**. Here given in the form: for every `n`, there exists a prime number `p ≥ n`. -/ theorem exists_infinite_primes (n : ℕ) : ∃ p, n ≤ p ∧ prime p := diff --git a/src/group_theory/order_of_element.lean b/src/group_theory/order_of_element.lean index 06b34f5afe778..7bf7c11411d5a 100644 --- a/src/group_theory/order_of_element.lean +++ b/src/group_theory/order_of_element.lean @@ -174,6 +174,10 @@ is_periodic_pt.minimal_period_dvd ((is_periodic_pt_mul_iff_pow_eq_one _).mpr h) lemma order_of_dvd_iff_pow_eq_one {n : ℕ} : order_of x ∣ n ↔ x ^ n = 1 := ⟨λ h, by rw [pow_eq_mod_order_of, nat.mod_eq_zero_of_dvd h, pow_zero], order_of_dvd_of_pow_eq_one⟩ +@[to_additive add_order_of_smul_dvd] +lemma order_of_pow_dvd (n : ℕ) : order_of (x ^ n) ∣ order_of x := +by rw [order_of_dvd_iff_pow_eq_one, pow_right_comm, pow_order_of_eq_one, one_pow] + @[to_additive add_order_of_map_dvd] lemma order_of_map_dvd {H : Type*} [monoid H] (ψ : G →* H) (x : G) : order_of (ψ x) ∣ order_of x := @@ -265,22 +269,36 @@ begin { rw [order_of_pow'' y m (hg.imp_symm order_of_eq_zero), h.gcd_eq_one, nat.div_one] }, end +namespace commute + +variables {x y} (h : commute x y) +include h + @[to_additive] -lemma commute.order_of_mul_dvd_lcm {x y : G} (h : commute x y) : - order_of (x * y) ∣ nat.lcm (order_of x) (order_of y) := +lemma order_of_mul_dvd_lcm : order_of (x * y) ∣ nat.lcm (order_of x) (order_of y) := begin convert function.commute.minimal_period_of_comp_dvd_lcm h.function_commute_mul_left, rw [order_of, comp_mul_left], end +@[to_additive] +lemma order_of_dvd_lcm_mul : order_of y ∣ nat.lcm (order_of x) (order_of (x * y)) := +begin + by_cases h0 : order_of x = 0, + { rw [h0, lcm_zero_left], apply dvd_zero }, + conv_lhs { rw [← one_mul y, ← pow_order_of_eq_one x, + ← succ_pred_eq_of_pos (nat.pos_of_ne_zero h0), pow_succ', mul_assoc] }, + exact (((commute.refl x).mul_right h).pow_left _).order_of_mul_dvd_lcm.trans + (lcm_dvd_iff.2 ⟨trans (order_of_pow_dvd _) (dvd_lcm_left _ _), dvd_lcm_right _ _⟩), +end + @[to_additive add_order_of_add_dvd_mul_add_order_of] -lemma commute.order_of_mul_dvd_mul_order_of {x y : G} (h : commute x y) : - order_of (x * y) ∣ (order_of x) * (order_of y) := +lemma order_of_mul_dvd_mul_order_of : order_of (x * y) ∣ (order_of x) * (order_of y) := dvd_trans h.order_of_mul_dvd_lcm (lcm_dvd_mul _ _) @[to_additive add_order_of_add_eq_mul_add_order_of_of_coprime] -lemma commute.order_of_mul_eq_mul_order_of_of_coprime {x y : G} (h : commute x y) - (hco : nat.coprime (order_of x) (order_of y)) : +lemma order_of_mul_eq_mul_order_of_of_coprime + (hco : (order_of x).coprime (order_of y)) : order_of (x * y) = (order_of x) * (order_of y) := begin convert h.function_commute_mul_left.minimal_period_of_comp_eq_mul_of_coprime hco, @@ -289,12 +307,34 @@ end /-- Commuting elements of finite order are closed under multiplication. -/ @[to_additive "Commuting elements of finite additive order are closed under addition."] -lemma commute.is_of_fin_order_mul - {x} (h : commute x y) (hx : is_of_fin_order x) (hy : is_of_fin_order y) : +lemma is_of_fin_order_mul + (hx : is_of_fin_order x) (hy : is_of_fin_order y) : is_of_fin_order (x * y) := order_of_pos_iff.mp $ pos_of_dvd_of_pos h.order_of_mul_dvd_mul_order_of $ mul_pos (order_of_pos' hx) (order_of_pos' hy) +/-- If each prime factor of `order_of x` has higher multiplicity in `order_of y`, and `x` commutes + with `y`, then the order of `x * y` is the same as the order of `y`. -/ +@[to_additive "If each prime factor of `add_order_of x` has higher multiplicity in `add_order_of y`, +and `x` commutes with `y`, then the order of `x * y` is the same as the order of `y`."] +lemma order_of_mul_eq_right_of_forall_prime_mul_dvd + (hy : is_of_fin_order y) + (hdvd : ∀ p : ℕ, p.prime → p ∣ order_of x → (p * order_of x) ∣ order_of y) : + order_of (x * y) = order_of y := +begin + have hoy := order_of_pos' hy, + have hxy := dvd_of_forall_prime_mul_dvd hdvd, + apply order_of_eq_of_pow_and_pow_div_prime hoy; simp only [ne, ← order_of_dvd_iff_pow_eq_one], + { exact trans h.order_of_mul_dvd_lcm (lcm_dvd hxy dvd_rfl) }, + refine λ p hp hpy hd, hp.ne_one _, + rw [← nat.dvd_one, ← mul_dvd_mul_iff_right hoy.ne', one_mul, ← dvd_div_iff hpy], + refine trans (order_of_dvd_lcm_mul h) (lcm_dvd_iff.2 ⟨(dvd_div_iff hpy).2 _, hd⟩), + by_cases p ∣ order_of x, + exacts [hdvd p hp h, (hp.coprime_iff_not_dvd.2 h).mul_dvd_of_dvd_of_dvd hpy hxy], +end + +end commute + section p_prime variables {a x n} {p : ℕ} [hp : fact p.prime] From a149cb73f18ddaf3573cc6ec74fe9f1c00e12497 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9mi=20Bottinelli?= Date: Thu, 22 Dec 2022 09:08:44 +0000 Subject: [PATCH 0097/1291] chore(analysis/bounded_variation): replace uses of `linarith` (#17998) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Rémi Bottinelli --- src/analysis/bounded_variation.lean | 54 +++++++++++++++++------------ 1 file changed, 31 insertions(+), 23 deletions(-) diff --git a/src/analysis/bounded_variation.lean b/src/analysis/bounded_variation.lean index f26499b3ab483..b8531f14bb883 100644 --- a/src/analysis/bounded_variation.lean +++ b/src/analysis/bounded_variation.lean @@ -238,17 +238,17 @@ begin { apply monotone_nat_of_le_succ (λ i, _), simp only [v], rcases lt_trichotomy i n with hi|rfl|hi, - { have : i + 1 ≤ n, by linarith, + { have : i + 1 ≤ n := nat.succ_le_of_lt hi, simp only [hi.le, this, if_true], exact hu (nat.le_succ i) }, { simp only [le_refl, if_true, add_le_iff_nonpos_right, le_zero_iff, nat.one_ne_zero, if_false, h], }, - { have A : ¬(i ≤ n), by linarith, - have B : ¬(i + 1 ≤ n), by linarith, + { have A : ¬(i ≤ n) := hi.not_le, + have B : ¬(i + 1 ≤ n) := λ h, A (i.le_succ.trans h), simp only [A, B, if_false]} }, refine ⟨v, n+2, hv, vs, (mem_image _ _ _).2 ⟨n+1, _, _⟩, _⟩, { rw mem_Iio, exact nat.lt_succ_self (n+1) }, - { have : ¬(n + 1 ≤ n), by linarith, + { have : ¬(n + 1 ≤ n) := nat.not_succ_le_self n, simp only [this, ite_eq_right_iff, is_empty.forall_iff] }, { calc ∑ i in finset.range n, edist (f (u (i+1))) (f (u i)) @@ -256,12 +256,12 @@ begin begin apply finset.sum_congr rfl (λ i hi, _), simp only [finset.mem_range] at hi, - have : i + 1 ≤ n, by linarith, + have : i + 1 ≤ n := nat.succ_le_of_lt hi, dsimp only [v], simp only [hi.le, this, if_true], end ... ≤ ∑ j in finset.range (n + 2), edist (f (v (j+1))) (f (v j)) : - finset.sum_le_sum_of_subset (finset.range_mono (by linarith)) } }, + finset.sum_le_sum_of_subset (finset.range_mono (nat.le_add_right n 2)) } }, have exists_N : ∃ N, N ≤ n ∧ x < u N, from ⟨n, le_rfl, h⟩, let N := nat.find exists_N, have hN : N ≤ n ∧ x < u N := nat.find_spec exists_N, @@ -275,22 +275,22 @@ begin { apply monotone_nat_of_le_succ (λ i, _), dsimp only [w], rcases lt_trichotomy (i + 1) N with hi|hi|hi, - { have : i < N, by linarith, + { have : i < N := nat.lt_of_le_of_lt (nat.le_succ i) hi, simp only [hi, this, if_true], exact hu (nat.le_succ _) }, - { have A : i < N, by linarith, - have B : ¬(i + 1 < N), by linarith, + { have A : i < N := hi ▸ (i.lt_succ_self), + have B : ¬(i + 1 < N) := by { rw ←hi, exact λ h, h.ne rfl, }, rw [if_pos A, if_neg B, if_pos hi], have T := nat.find_min exists_N A, push_neg at T, exact T (A.le.trans hN.1) }, - { have A : ¬(i < N), by linarith, - have B : ¬(i + 1 < N), by linarith, - have C : ¬(i + 1 = N), by linarith, + { have A : ¬(i < N) := (nat.lt_succ_iff.mp hi).not_lt, + have B : ¬(i + 1 < N) := hi.not_lt, + have C : ¬(i + 1 = N) := hi.ne.symm, have D : i + 1 - 1 = i := nat.pred_succ i, rw [if_neg A, if_neg B, if_neg C, D], split_ifs, - { exact hN.2.le.trans (hu (by linarith)) }, + { exact hN.2.le.trans (hu (le_of_not_lt A)) }, { exact hu (nat.pred_le _) } } }, refine ⟨w, n+1, hw, ws, (mem_image _ _ _).2 ⟨N, hN.1.trans_lt (nat.lt_succ_self n), _⟩, _⟩, { dsimp only [w], rw [if_neg (lt_irrefl N), if_pos rfl] }, @@ -335,7 +335,7 @@ begin simp only [finset.mem_Ico, zero_le', true_and] at hi, dsimp only [w], have A : i + 1 < N, from nat.lt_pred_iff.1 hi, - have B : i < N, by linarith, + have B : i < N := nat.lt_of_succ_lt A, rw [if_pos A, if_pos B] }, { have A : N - 1 + 1 = N, from nat.succ_pred_eq_of_pos Npos, have : finset.Ico (N - 1) N = {N - 1}, by rw [← nat.Ico_succ_singleton, A], @@ -343,13 +343,19 @@ begin { apply finset.sum_congr rfl (λ i hi, _), simp only [finset.mem_Ico] at hi, dsimp only [w], - have A : ¬(1 + i + 1 < N), by linarith, - have B : ¬(1 + i + 1 = N), by linarith, - have C : ¬(1 + i < N), by linarith, - have D : ¬(1 + i = N), by linarith, + have A : ¬(1 + i + 1 < N) := λ h, by + { rw [add_assoc, add_comm] at h, + exact (hi.left).not_lt ((i.lt_succ_self).trans ((i.succ.lt_succ_self).trans h)), }, + have B : ¬(1 + i + 1 = N) := λ h, by + { rw [←h, add_assoc, add_comm] at hi, + exact nat.not_succ_le_self i (i.succ.le_succ.trans hi.left), }, + have C : ¬(1 + i < N) := λ h, by + { rw [add_comm] at h, exact (hi.left).not_lt (i.lt_succ_self.trans h), }, + have D : ¬(1 + i = N) := λ h, by + { rw [←h, add_comm, nat.succ_le_iff] at hi, exact hi.left.ne rfl, }, rw [if_neg A, if_neg B, if_neg C, if_neg D], congr' 3; - { rw eq_tsub_iff_add_eq_of_le, { abel }, { linarith } } } + { rw [add_comm, nat.sub_one], apply nat.pred_succ, } } end ... = ∑ i in finset.Ico 0 (N-1), edist (f (w (i + 1))) (f (w i)) + edist (f (w (N + 1))) (f (w (N - 1))) + @@ -357,7 +363,7 @@ begin begin congr' 1, congr' 1, { dsimp only [w], - have A : ¬(N + 1 < N), by linarith, + have A : ¬(N + 1 < N) := nat.not_succ_lt_self, have B : N - 1 < N := nat.pred_lt Npos.ne', simp only [A, not_and, not_lt, nat.succ_ne_self, nat.add_succ_sub_one, add_zero, if_false, B, if_true] }, @@ -368,9 +374,9 @@ begin ∑ i in finset.Ico (N + 1) (n + 1), edist (f (w (i + 1))) (f (w i)) : begin refine add_le_add (add_le_add le_rfl _) le_rfl, - have A : N - 1 + 1 = N, from nat.succ_pred_eq_of_pos Npos, - have B : N - 1 + 1 < N + 1, by linarith, - have C : N - 1 < N + 1, by linarith, + have A : N - 1 + 1 = N := nat.succ_pred_eq_of_pos Npos, + have B : N - 1 + 1 < N + 1 := A.symm ▸ N.lt_succ_self, + have C : N - 1 < N + 1 := lt_of_le_of_lt (N.pred_le) (N.lt_succ_self), rw [finset.sum_eq_sum_Ico_succ_bot C, finset.sum_eq_sum_Ico_succ_bot B, A, finset.Ico_self, finset.sum_empty, add_zero, add_comm (edist _ _)], exact edist_triangle _ _ _, @@ -750,3 +756,5 @@ lemma lipschitz_with.ae_differentiable_at {C : ℝ≥0} {f : ℝ → V} (h : lipschitz_with C f) : ∀ᵐ x, differentiable_at ℝ f x := (h.has_locally_bounded_variation_on univ).ae_differentiable_at + + From 2f3994e1b117b1e1da49bcfb67334f33460c3ce4 Mon Sep 17 00:00:00 2001 From: Gabriel Ebner Date: Thu, 22 Dec 2022 11:30:20 +0000 Subject: [PATCH 0098/1291] feat(*): bump to lean 3.50.1 (#17995) --- leanpkg.toml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/leanpkg.toml b/leanpkg.toml index e760ceea83ade..c50815f219262 100644 --- a/leanpkg.toml +++ b/leanpkg.toml @@ -1,7 +1,7 @@ [package] name = "mathlib" version = "0.1" -lean_version = "leanprover-community/lean:3.49.1" +lean_version = "leanprover-community/lean:3.50.1" path = "src" [dependencies] From b26e15a46f1a713ce7410e016d50575bb0bc3aa4 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 22 Dec 2022 19:09:30 +0000 Subject: [PATCH 0099/1291] refactor(ring_theory/ideal/quotient): extract a `ring_con` structure (#17833) The intent here is to remove some duplication with `ring_quot`. I've only copied across the basic lemmas from `group_theory/congruence`; I imagine most of the rest won't be that useful, and that if we do want them, we should look for a means to deduplicate them. Note that `ring_con.quotient` can be used to take quotients of semirings. --- src/ring_theory/congruence.lean | 271 ++++++++++++++++++++++++++++ src/ring_theory/ideal/quotient.lean | 42 ++--- 2 files changed, 285 insertions(+), 28 deletions(-) create mode 100644 src/ring_theory/congruence.lean diff --git a/src/ring_theory/congruence.lean b/src/ring_theory/congruence.lean new file mode 100644 index 0000000000000..a79ed4a5a51b9 --- /dev/null +++ b/src/ring_theory/congruence.lean @@ -0,0 +1,271 @@ +/- +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.group_ring_action.basic +import algebra.hom.ring +import algebra.ring.inj_surj +import group_theory.congruence + +/-! +# Congruence relations on rings + +This file defines congruence relations on rings, which extend `con` and `add_con` on monoids and +additive monoids. + +Most of the time you likely want to use the `ideal.quotient` API that is built on top of this. + +## Main Definitions + +* `ring_con R`: the type of congruence relations respecting `+` and `*`. +* `ring_con_gen r`: the inductively defined smallest ring congruence relation containing a given + binary relation. + +## TODO + +* Use this for `ring_quot` too. +* Copy across more API from `con` and `add_con` in `group_theory/congruence.lean`, such as: + * The `complete_lattice` structure. + * The `con_gen_eq` lemma, stating that + `ring_con_gen r = Inf {s : ring_con M | ∀ x y, r x y → s x y}`. +-/ + +/-- A congruence relation on a type with an addition and multiplication is an equivalence relation +which preserves both. -/ +/- Note: we can't extend both `add_con R` and `mul_con R` in Lean 3 due to interactions between old- +and new-style structures. We can revisit this in Lean 4. (After and not during the port!) -/ +structure ring_con (R : Type*) [has_add R] [has_mul R] extends setoid R := +(add' : ∀ {w x y z}, r w x → r y z → r (w + y) (x + z)) +(mul' : ∀ {w x y z}, r w x → r y z → r (w * y) (x * z)) + +variables {α R : Type*} + +/-- The inductively defined smallest ring congruence relation containing a given binary + relation. -/ +inductive ring_con_gen.rel [has_add R] [has_mul R] (r : R → R → Prop) : R → R → Prop +| of : Π x y, r x y → ring_con_gen.rel x y +| refl : Π x, ring_con_gen.rel x x +| symm : Π {x y}, ring_con_gen.rel x y → ring_con_gen.rel y x +| trans : Π {x y z}, ring_con_gen.rel x y → ring_con_gen.rel y z → ring_con_gen.rel x z +| add : Π {w x y z}, ring_con_gen.rel w x → ring_con_gen.rel y z → ring_con_gen.rel (w + y) (x + z) +| mul : Π {w x y z}, ring_con_gen.rel w x → ring_con_gen.rel y z → ring_con_gen.rel (w * y) (x * z) + +/-- The inductively defined smallest ring congruence relation containing a given binary + relation. -/ +def ring_con_gen [has_add R] [has_mul R] (r : R → R → Prop) : ring_con R := +{ r := ring_con_gen.rel r, + iseqv := ⟨ring_con_gen.rel.refl, @ring_con_gen.rel.symm _ _ _ _, @ring_con_gen.rel.trans _ _ _ _⟩, + add' := λ _ _ _ _, ring_con_gen.rel.add, + mul' := λ _ _ _ _, ring_con_gen.rel.mul } + +namespace ring_con + +section basic +variables [has_add R] [has_mul R] (c : ring_con R) + +/-- Every `ring_con` is also an `add_con` -/ +def to_add_con : add_con R := { ..c } + +/-- Every `ring_con` is also a `con` -/ +def to_con : con R := { ..c } + +/-- A coercion from a congruence relation to its underlying binary relation. -/ +instance : has_coe_to_fun (ring_con R) (λ _, R → R → Prop) := ⟨λ c, c.r⟩ + +@[simp] lemma rel_eq_coe : c.r = c := rfl + +protected lemma refl (x) : c x x := c.refl' x +protected lemma symm {x y} : c x y → c y x := c.symm' +protected lemma trans {x y z} : c x y → c y z → c x z := c.trans' +protected lemma add {w x y z} : c w x → c y z → c (w + y) (x + z) := c.add' +protected lemma mul {w x y z} : c w x → c y z → c (w * y) (x * z) := c.mul' + +@[simp] lemma rel_mk {s : setoid R} {ha hm a b} : ring_con.mk s ha hm a b ↔ setoid.r a b := iff.rfl + +instance : inhabited (ring_con R) := ⟨ring_con_gen empty_relation⟩ + +end basic + +section quotient + +section basic +variables [has_add R] [has_mul R] (c : ring_con R) +/-- Defining the quotient by a congruence relation of a type with addition and multiplication. -/ +protected def quotient := quotient c.to_setoid + +/-- Coercion from a type with addition and multiplication to its quotient by a congruence relation. + +See Note [use has_coe_t]. -/ +instance : has_coe_t R c.quotient := ⟨@quotient.mk _ c.to_setoid⟩ + +/-- The quotient by a decidable congruence relation has decidable equality. -/ +-- Lower the priority since it unifies with any quotient type. +@[priority 500] instance [d : ∀ a b, decidable (c a b)] : decidable_eq c.quotient := +@quotient.decidable_eq R c.to_setoid d + +@[simp] lemma quot_mk_eq_coe (x : R) : quot.mk c x = (x : c.quotient) := rfl + +/-- Two elements are related by a congruence relation `c` iff they are represented by the same +element of the quotient by `c`. -/ +@[simp] protected lemma eq {a b : R} : (a : c.quotient) = b ↔ c a b := quotient.eq' + +end basic + +/-! ### Basic notation + +The basic algebraic notation, `0`, `1`, `+`, `*`, `-`, `^`, descend naturally under the quotient +-/ +section data + +section add_mul +variables [has_add R] [has_mul R] (c : ring_con R) +instance : has_add c.quotient := c.to_add_con.has_add +@[simp, norm_cast] lemma coe_add (x y : R) : (↑(x + y) : c.quotient) = ↑x + ↑y := rfl +instance : has_mul c.quotient := c.to_con.has_mul +@[simp, norm_cast] lemma coe_mul (x y : R) : (↑(x * y) : c.quotient) = ↑x * ↑y := rfl +end add_mul + +section zero +variables [add_zero_class R] [has_mul R] (c : ring_con R) +instance : has_zero c.quotient := c.to_add_con^.quotient.has_zero +@[simp, norm_cast] lemma coe_zero : (↑(0 : R) : c.quotient) = 0 := rfl +end zero + +section one +variables [has_add R] [mul_one_class R] (c : ring_con R) +instance : has_one c.quotient := c.to_con^.quotient.has_one +@[simp, norm_cast] lemma coe_one : (↑(1 : R) : c.quotient) = 1 := rfl +end one + +section smul +variables [has_add R] [mul_one_class R] [has_smul α R] [is_scalar_tower α R R] (c : ring_con R) +instance : has_smul α c.quotient := c.to_con.has_smul +@[simp, norm_cast] lemma coe_smul (a : α) (x : R) : (↑(a • x) : c.quotient) = a • x := rfl +end smul + +section neg_sub_zsmul +variables [add_group R] [has_mul R] (c : ring_con R) +instance : has_neg c.quotient := c.to_add_con^.has_neg +@[simp, norm_cast] lemma coe_neg (x : R) : (↑(-x) : c.quotient) = -x := rfl +instance : has_sub c.quotient := c.to_add_con^.has_sub +@[simp, norm_cast] lemma coe_sub (x y : R) : (↑(x - y) : c.quotient) = x - y := rfl +instance has_zsmul : has_smul ℤ c.quotient := c.to_add_con^.quotient.has_zsmul +@[simp, norm_cast] lemma coe_zsmul (z : ℤ) (x : R) : (↑(z • x) : c.quotient) = z • x := rfl +end neg_sub_zsmul + +section nsmul +variables [add_monoid R] [has_mul R] (c : ring_con R) +instance has_nsmul : has_smul ℕ c.quotient := c.to_add_con^.quotient.has_nsmul +@[simp, norm_cast] lemma coe_nsmul (n : ℕ) (x : R) : (↑(n • x) : c.quotient) = n • x := rfl +end nsmul + +section pow +variables [has_add R] [monoid R] (c : ring_con R) +instance : has_pow c.quotient ℕ := c.to_con^.nat.has_pow +@[simp, norm_cast] lemma coe_pow (x : R) (n : ℕ) : (↑(x ^ n) : c.quotient) = x ^ n := rfl +end pow + +section nat_cast +variables [add_monoid_with_one R] [has_mul R] (c : ring_con R) +instance : has_nat_cast c.quotient := ⟨λ n, ↑(n : R)⟩ +@[simp, norm_cast] lemma coe_nat_cast (n : ℕ) : (↑(n : R) : c.quotient) = n := rfl +end nat_cast + +section int_cast +variables [add_group_with_one R] [has_mul R] (c : ring_con R) +instance : has_int_cast c.quotient := ⟨λ z, ↑(z : R)⟩ +@[simp, norm_cast] lemma coe_int_cast (n : ℕ) : (↑(n : R) : c.quotient) = n := rfl +end int_cast + +instance [inhabited R] [has_add R] [has_mul R] (c : ring_con R) : inhabited c.quotient := +⟨↑(default : R)⟩ + +end data + +/-! ### Algebraic structure + +The operations above on the quotient by `c : ring_con R` preseverse the algebraic structure of `R`. +-/ + +section algebraic + +instance [non_unital_non_assoc_semiring R] (c : ring_con R) : + non_unital_non_assoc_semiring c.quotient := +function.surjective.non_unital_non_assoc_semiring _ quotient.surjective_quotient_mk' + rfl (λ _ _, rfl) (λ _ _, rfl) (λ _ _, rfl) + +instance [non_assoc_semiring R] (c : ring_con R) : + non_assoc_semiring c.quotient := +function.surjective.non_assoc_semiring _ quotient.surjective_quotient_mk' + rfl rfl (λ _ _, rfl) (λ _ _, rfl) (λ _ _, rfl) (λ _, rfl) + +instance [non_unital_semiring R] (c : ring_con R) : + non_unital_semiring c.quotient := +function.surjective.non_unital_semiring _ quotient.surjective_quotient_mk' + rfl (λ _ _, rfl) (λ _ _, rfl) (λ _ _, rfl) + +instance [semiring R] (c : ring_con R) : + semiring c.quotient := +function.surjective.semiring _ quotient.surjective_quotient_mk' + rfl rfl (λ _ _, rfl) (λ _ _, rfl) (λ _ _, rfl) (λ _ _, rfl) (λ _, rfl) + +instance [comm_semiring R] (c : ring_con R) : + comm_semiring c.quotient := +function.surjective.comm_semiring _ quotient.surjective_quotient_mk' + rfl rfl (λ _ _, rfl) (λ _ _, rfl) (λ _ _, rfl) (λ _ _, rfl) (λ _, rfl) + +instance [non_unital_non_assoc_ring R] (c : ring_con R) : + non_unital_non_assoc_ring c.quotient := +function.surjective.non_unital_non_assoc_ring _ quotient.surjective_quotient_mk' + rfl (λ _ _, rfl) (λ _ _, rfl) (λ _, rfl) (λ _ _, rfl) (λ _ _, rfl) (λ _ _, rfl) + +instance [non_assoc_ring R] (c : ring_con R) : + non_assoc_ring c.quotient := +function.surjective.non_assoc_ring _ quotient.surjective_quotient_mk' + rfl rfl (λ _ _, rfl) (λ _ _, rfl) (λ _, rfl) (λ _ _, rfl) (λ _ _, rfl) (λ _ _, rfl) (λ _, rfl) + (λ _, rfl) + +instance [non_unital_ring R] (c : ring_con R) : + non_unital_ring c.quotient := +function.surjective.non_unital_ring _ quotient.surjective_quotient_mk' + rfl (λ _ _, rfl) (λ _ _, rfl) (λ _, rfl) (λ _ _, rfl) (λ _ _, rfl) (λ _ _, rfl) + +instance [ring R] (c : ring_con R) : + ring c.quotient := +function.surjective.ring _ quotient.surjective_quotient_mk' + rfl rfl (λ _ _, rfl) (λ _ _, rfl) (λ _, rfl) (λ _ _, rfl) (λ _ _, rfl) (λ _ _, rfl) (λ _ _, rfl) + (λ _, rfl) (λ _, rfl) + +instance [comm_ring R] (c : ring_con R) : + comm_ring c.quotient := +function.surjective.comm_ring _ quotient.surjective_quotient_mk' + rfl rfl (λ _ _, rfl) (λ _ _, rfl) (λ _, rfl) (λ _ _, rfl) (λ _ _, rfl) (λ _ _, rfl) (λ _ _, rfl) + (λ _, rfl) (λ _, rfl) + +instance [monoid α] [non_assoc_semiring R] [distrib_mul_action α R] [is_scalar_tower α R R] + (c : ring_con R) : + distrib_mul_action α c.quotient := +{ smul := (•), + smul_zero := λ r, congr_arg quotient.mk' $ smul_zero _, + smul_add := λ r, quotient.ind₂' $ by exact λ m₁ m₂, congr_arg quotient.mk' $ smul_add _ _ _, + .. c.to_con.mul_action } + +instance [monoid α] [semiring R] [mul_semiring_action α R] [is_scalar_tower α R R] + (c : ring_con R) : + mul_semiring_action α c.quotient := +{ smul := (•), + .. c^.quotient.distrib_mul_action, + .. c.to_con.mul_distrib_mul_action } + +end algebraic + +/-- The natural homomorphism from a ring to its quotient by a congruence relation. -/ +def mk' [non_assoc_semiring R] (c : ring_con R) : R →+* c.quotient := +{ to_fun := quotient.mk', map_zero' := rfl, map_one' := rfl, + map_add' := λ _ _, rfl, map_mul' := λ _ _, rfl } + +end quotient + +end ring_con diff --git a/src/ring_theory/ideal/quotient.lean b/src/ring_theory/ideal/quotient.lean index 4ddab13391a29..a1b0740ec5350 100644 --- a/src/ring_theory/ideal/quotient.lean +++ b/src/ring_theory/ideal/quotient.lean @@ -5,6 +5,7 @@ Authors: Kenny Lau, Chris Hughes, Mario Carneiro, Anne Baanen -/ import algebra.ring.fin import linear_algebra.quotient +import ring_theory.congruence import ring_theory.ideal.basic import tactic.fin_cases /-! @@ -51,36 +52,21 @@ variables {I} {x y : R} instance has_one (I : ideal R) : has_one (R ⧸ I) := ⟨submodule.quotient.mk 1⟩ -instance has_mul (I : ideal R) : has_mul (R ⧸ I) := -⟨λ a b, quotient.lift_on₂' a b (λ a b, submodule.quotient.mk (a * b)) $ - λ a₁ a₂ b₁ b₂ h₁ h₂, quot.sound $ begin - rw submodule.quotient_rel_r_def at h₁ h₂ ⊢, - have F := I.add_mem (I.mul_mem_left a₂ h₁) (I.mul_mem_right b₁ h₂), - have : a₁ * a₂ - b₁ * b₂ = a₂ * (a₁ - b₁) + (a₂ - b₂) * b₁, - { rw [mul_sub, sub_mul, sub_add_sub_cancel, mul_comm, mul_comm b₁] }, - rw ← this at F, - change _ ∈ _, convert F, -end⟩ +/-- On `ideal`s, `submodule.quotient_rel` is a ring congruence. -/ +protected def ring_con (I : ideal R) : ring_con R := +{ mul' := λ a₁ b₁ a₂ b₂ h₁ h₂, begin + rw submodule.quotient_rel_r_def at h₁ h₂ ⊢, + have F := I.add_mem (I.mul_mem_left a₂ h₁) (I.mul_mem_right b₁ h₂), + have : a₁ * a₂ - b₁ * b₂ = a₂ * (a₁ - b₁) + (a₂ - b₂) * b₁, + { rw [mul_sub, sub_mul, sub_add_sub_cancel, mul_comm, mul_comm b₁] }, + rw ← this at F, + change _ ∈ _, convert F, + end, + .. quotient_add_group.con I.to_add_subgroup } instance comm_ring (I : ideal R) : comm_ring (R ⧸ I) := -{ mul := (*), - one := 1, - nat_cast := λ n, submodule.quotient.mk n, - nat_cast_zero := by simp [nat.cast], - nat_cast_succ := by simp [nat.cast]; refl, - mul_assoc := λ a b c, quotient.induction_on₃' a b c $ - λ a b c, congr_arg submodule.quotient.mk (mul_assoc a b c), - mul_comm := λ a b, quotient.induction_on₂' a b $ - λ a b, congr_arg submodule.quotient.mk (mul_comm a b), - one_mul := λ a, quotient.induction_on' a $ - λ a, congr_arg submodule.quotient.mk (one_mul a), - mul_one := λ a, quotient.induction_on' a $ - λ a, congr_arg submodule.quotient.mk (mul_one a), - left_distrib := λ a b c, quotient.induction_on₃' a b c $ - λ a b c, congr_arg submodule.quotient.mk (left_distrib a b c), - right_distrib := λ a b c, quotient.induction_on₃' a b c $ - λ a b c, congr_arg submodule.quotient.mk (right_distrib a b c), - ..submodule.quotient.add_comm_group I } +{ ..submodule.quotient.add_comm_group I, -- to help with unification + ..(quotient.ring_con I)^.quotient.comm_ring } /-- The ring homomorphism from a ring `R` to a quotient ring `R/I`. -/ def mk (I : ideal R) : R →+* (R ⧸ I) := From 51c6beb41be93b7b4fad829e1c74b9c3939e273d Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Fri, 23 Dec 2022 08:51:57 +0000 Subject: [PATCH 0100/1291] feat(number_theory/number_field/embeddings) : add infinite_places (#17844) This PR adds definitions and basic lemmas about places, complex embeddings and infinite places of number field. --- src/algebra/star/basic.lean | 6 +- src/analysis/normed/field/basic.lean | 6 + .../number_field/embeddings.lean | 136 +++++++++++++++++- src/topology/algebra/ring.lean | 15 ++ src/topology/metric_space/cau_seq_filter.lean | 8 +- 5 files changed, 155 insertions(+), 16 deletions(-) diff --git a/src/algebra/star/basic.lean b/src/algebra/star/basic.lean index 2229212f9d368..b9e18e02a0c72 100644 --- a/src/algebra/star/basic.lean +++ b/src/algebra/star/basic.lean @@ -291,16 +291,16 @@ lemma star_ring_end_apply [comm_semiring R] [star_ring R] {x : R} : @[simp] lemma star_ring_end_self_apply [comm_semiring R] [star_ring R] (x : R) : star_ring_end R (star_ring_end R x) = x := star_star x -instance ring_hom.has_involutive_star {S : Type u} [non_assoc_semiring S] [comm_semiring R] +instance ring_hom.has_involutive_star {S : Type*} [non_assoc_semiring S] [comm_semiring R] [star_ring R] : has_involutive_star (S →+* R) := { to_has_star := { star := λ f, ring_hom.comp (star_ring_end R) f }, star_involutive := by { intro _, ext, simp only [ring_hom.coe_comp, function.comp_app, star_ring_end_self_apply] }} -lemma ring_hom.star_def {S : Type u} [non_assoc_semiring S] [comm_semiring R] [star_ring R] +lemma ring_hom.star_def {S : Type*} [non_assoc_semiring S] [comm_semiring R] [star_ring R] (f : S →+* R) : has_star.star f = ring_hom.comp (star_ring_end R) f := rfl -lemma ring_hom.star_apply {S : Type u} [non_assoc_semiring S] [comm_semiring R] [star_ring R] +lemma ring_hom.star_apply {S : Type*} [non_assoc_semiring S] [comm_semiring R] [star_ring R] (f : S →+* R) (s : S) : star f s = star (f s) := rfl -- A more convenient name for complex conjugation diff --git a/src/analysis/normed/field/basic.lean b/src/analysis/normed/field/basic.lean index 15dd3438859d1..ae8b07693a29b 100644 --- a/src/analysis/normed/field/basic.lean +++ b/src/analysis/normed/field/basic.lean @@ -414,6 +414,12 @@ 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]⟩ +instance is_absolute_value_norm : is_absolute_value (norm : α → ℝ) := +{ abv_nonneg := norm_nonneg, + abv_eq_zero := λ _, norm_eq_zero, + abv_add := norm_add_le, + abv_mul := norm_mul } + @[simp] lemma nnnorm_mul (a b : α) : ‖a * b‖₊ = ‖a‖₊ * ‖b‖₊ := nnreal.eq $ norm_mul a b diff --git a/src/number_theory/number_field/embeddings.lean b/src/number_theory/number_field/embeddings.lean index 310babad710b7..23616031e974f 100644 --- a/src/number_theory/number_field/embeddings.lean +++ b/src/number_theory/number_field/embeddings.lean @@ -4,25 +4,27 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Alex J. Best, Xavier Roblot -/ -import analysis.normed_space.basic import number_theory.number_field.basic -import topology.algebra.polynomial +import analysis.complex.polynomial +import data.complex.basic /-! # Embeddings of number fields This file defines the embeddings of a number field into an algebraic closed field. ## Main Results -* `number_field.embeddings.eq_roots`: let `x ∈ K` with `K` number field and let `A` be an algebraic - closed field of char. 0, then the images of `x` by the embeddings of `K` in `A` are exactly the - roots in `A` of the minimal polynomial of `x` over `ℚ`. +* `number_field.embeddings.range_eval_eq_root_set_minpoly`: let `x ∈ K` with `K` number field and + let `A` be an algebraic closed field of char. 0, then the images of `x` by the embeddings of `K` + in `A` are exactly the roots in `A` of the minimal polynomial of `x` over `ℚ`. * `number_field.embeddings.pow_eq_one_of_norm_eq_one`: an algebraic integer whose conjugates are all of norm one is a root of unity. ## Tags -number field, embeddings +number field, embeddings, places, infinite places -/ +open_locale classical + namespace number_field.embeddings section fintype @@ -42,6 +44,12 @@ variables [is_alg_closed A] lemma card : fintype.card (K →+* A) = finrank ℚ K := by rw [fintype.of_equiv_card ring_hom.equiv_rat_alg_hom.symm, alg_hom.card] +instance : nonempty (K →+* A) := +begin + rw [← fintype.card_pos_iff, number_field.embeddings.card K A], + exact finite_dimensional.finrank_pos, +end + end fintype section roots @@ -123,3 +131,119 @@ end end bounded end number_field.embeddings + +section place + +variables {K : Type*} [field K] {A : Type*} [normed_division_ring A] [nontrivial A] (φ : K →+* A) + +/-- An embedding into a normed division ring defines a place of `K` -/ +def number_field.place : absolute_value K ℝ := +(is_absolute_value.to_absolute_value (norm : A → ℝ)).comp φ.injective + +@[simp] +lemma number_field.place_apply (x : K) : (number_field.place φ) x = norm (φ x) := rfl + +end place + +namespace number_field.complex_embedding + +open complex number_field + +open_locale complex_conjugate + +variables {K : Type*} [field K] + +/-- The conjugate of a complex embedding as a complex embedding. -/ +@[reducible] def conjugate (φ : K →+* ℂ) : K →+* ℂ := star φ + +@[simp] +lemma conjugate_coe_eq (φ : K →+* ℂ) (x : K) : (conjugate φ) x = conj (φ x) := rfl + +lemma place_conjugate (φ : K →+* ℂ) : place (conjugate φ) = place φ := +by { ext, simp only [place_apply, norm_eq_abs, abs_conj, conjugate_coe_eq] } + +/-- A embedding into `ℂ` is real if it is fixed by complex conjugation. -/ +@[reducible] def is_real (φ : K →+* ℂ) : Prop := is_self_adjoint φ + +lemma is_real_iff {φ : K →+* ℂ} : is_real φ ↔ conjugate φ = φ := is_self_adjoint_iff + +/-- A real embedding as a ring homomorphism from `K` to `ℝ` . -/ +def is_real.embedding {φ : K →+* ℂ} (hφ : is_real φ) : K →+* ℝ := +{ to_fun := λ x, (φ x).re, + map_one' := by simp only [map_one, one_re], + map_mul' := by simp only [complex.eq_conj_iff_im.mp (ring_hom.congr_fun hφ _), map_mul, mul_re, + mul_zero, tsub_zero, eq_self_iff_true, forall_const], + map_zero' := by simp only [map_zero, zero_re], + map_add' := by simp only [map_add, add_re, eq_self_iff_true, forall_const], } + +@[simp] +lemma is_real.coe_embedding_apply {φ : K →+* ℂ} (hφ : is_real φ) (x : K) : + (hφ.embedding x : ℂ) = φ x := +begin + ext, { refl, }, + { rw [of_real_im, eq_comm, ← complex.eq_conj_iff_im], + rw is_real at hφ, + exact ring_hom.congr_fun hφ x, }, +end + +lemma is_real.place_embedding {φ : K →+* ℂ} (hφ : is_real φ) : + place hφ.embedding = place φ := +by { ext x, simp only [place_apply, real.norm_eq_abs, ←abs_of_real, norm_eq_abs, + hφ.coe_embedding_apply x], } + +lemma is_real_conjugate_iff {φ : K →+* ℂ} : + is_real (conjugate φ) ↔ is_real φ := is_self_adjoint.star_iff + +end number_field.complex_embedding + +section infinite_place + +open number_field + +variables (K : Type*) [field K] + +/-- An infinite place of a number field `K` is a place associated to a complex embedding. -/ +def number_field.infinite_place := { w : absolute_value K ℝ // ∃ φ : K →+* ℂ, place φ = w} + +instance [number_field K] : nonempty (number_field.infinite_place K) := set.range.nonempty _ + +variables {K} + +/-- Return the infinite place defined by a complex embedding `φ`. -/ +noncomputable def number_field.infinite_place.mk (φ : K →+* ℂ) : number_field.infinite_place K := +⟨place φ, ⟨φ, rfl⟩⟩ + +namespace number_field.infinite_place + +open number_field + +instance : has_coe_to_fun (infinite_place K) (λ _, K → ℝ) := { coe := λ w, w.1 } + +instance : monoid_with_zero_hom_class (infinite_place K) K ℝ := +{ coe := λ w x, w.1 x, + coe_injective' := λ _ _ h, subtype.eq (absolute_value.ext (λ x, congr_fun h x)), + map_mul := λ w _ _, w.1.map_mul _ _, + map_one := λ w, w.1.map_one, + map_zero := λ w, w.1.map_zero, } + +instance : nonneg_hom_class (infinite_place K) K ℝ := +{ coe := λ w x, w x, + coe_injective' := λ _ _ h, subtype.eq (absolute_value.ext (λ x, congr_fun h x)), + map_nonneg := λ w x, w.1.nonneg _ } + +lemma coe_mk (φ : K →+* ℂ) : ⇑(mk φ) = place φ := rfl + +lemma apply (φ : K →+* ℂ) (x : K) : (mk φ) x = complex.abs (φ x) := rfl + +/-- For an infinite place `w`, return an embedding `φ` such that `w = infinite_place φ` . -/ +noncomputable def embedding (w : infinite_place K) : K →+* ℂ := (w.2).some + +lemma mk_embedding (w : infinite_place K) : + mk (embedding w) = w := +subtype.ext (w.2).some_spec + +lemma pos_iff (w : infinite_place K) (x : K) : 0 < w x ↔ x ≠ 0 := absolute_value.pos_iff w.1 + +end number_field.infinite_place + +end infinite_place diff --git a/src/topology/algebra/ring.lean b/src/topology/algebra/ring.lean index 9ee8148180bc4..8b5d89f91bb39 100644 --- a/src/topology/algebra/ring.lean +++ b/src/topology/algebra/ring.lean @@ -472,3 +472,18 @@ def to_add_group_topology.order_embedding : order_embedding (ring_topology α) end } end ring_topology + +section absolute_value + +/-- Construct an absolute value on a semiring `T` from an absolute value on a semiring `R` +and an injective ring homomorphism `f : T →+* R` -/ +def absolute_value.comp {R S T : Type*} [semiring T] [semiring R] [ordered_semiring S] + (v : absolute_value R S) {f : T →+* R} (hf : function.injective f) : + absolute_value T S := +{ to_fun := v ∘ f, + map_mul' := by simp only [function.comp_app, map_mul, eq_self_iff_true, forall_const], + nonneg' := by simp only [v.nonneg, forall_const], + eq_zero' := by simp only [map_eq_zero_iff f hf, v.eq_zero, forall_const, iff_self], + add_le' := by simp only [function.comp_app, map_add, v.add_le, forall_const], } + +end absolute_value diff --git a/src/topology/metric_space/cau_seq_filter.lean b/src/topology/metric_space/cau_seq_filter.lean index db392d2b56dcb..ac6401bc72230 100644 --- a/src/topology/metric_space/cau_seq_filter.lean +++ b/src/topology/metric_space/cau_seq_filter.lean @@ -41,16 +41,10 @@ variables [normed_field β] This section shows that if we have a uniform space generated by an absolute value, topological completeness and Cauchy sequence completeness coincide. The problem is that there isn't a good notion of "uniform space generated by an absolute value", so right now this is - specific to norm. Furthermore, norm only instantiates is_absolute_value on normed_field. + specific to norm. Furthermore, norm only instantiates is_absolute_value on normed_division_ring. This needs to be fixed, since it prevents showing that ℤ_[hp] is complete -/ -instance normed_field.is_absolute_value : is_absolute_value (norm : β → ℝ) := -{ abv_nonneg := norm_nonneg, - abv_eq_zero := λ _, norm_eq_zero, - abv_add := norm_add_le, - abv_mul := norm_mul } - open metric lemma cauchy_seq.is_cau_seq {f : ℕ → β} (hf : cauchy_seq f) : From cfa57f46c99598023cb2dcb19d2484024e496ba6 Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Fri, 23 Dec 2022 11:28:41 +0000 Subject: [PATCH 0101/1291] chore(*): fix docs and a name generated by to_additive (#18002) The autogenerated name was `order_of_add_eq_right_of_forall_prime_add_dvd`, but the second `add` shouldn't be to_additivized. Now corrected to `add_order_of_add_eq_right_of_forall_prime_mul_dvd`. I've checked the namespace still gets correctly additivized to `add_commute`. Co-authored-by: Oliver Nash Co-authored-by: Junyan Xu --- src/geometry/manifold/mfderiv.lean | 2 +- src/group_theory/order_of_element.lean | 9 +++++---- src/tactic/print_sorry.lean | 4 ++-- 3 files changed, 8 insertions(+), 7 deletions(-) diff --git a/src/geometry/manifold/mfderiv.lean b/src/geometry/manifold/mfderiv.lean index 1dbd812eeb894..526119e511de0 100644 --- a/src/geometry/manifold/mfderiv.lean +++ b/src/geometry/manifold/mfderiv.lean @@ -1284,7 +1284,7 @@ lemma has_mfderiv_at.neg (hf : has_mfderiv_at I 𝓘(𝕜, E') f z f') : lemma mdifferentiable_at.neg (hf : mdifferentiable_at I 𝓘(𝕜, E') f z) : mdifferentiable_at I 𝓘(𝕜, E') (-f) z := -(hf.has_mfderiv_at.neg ).mdifferentiable_at +hf.has_mfderiv_at.neg.mdifferentiable_at lemma mdifferentiable.neg {f : M → E'} (hf : mdifferentiable I 𝓘(𝕜, E') f) : mdifferentiable I 𝓘(𝕜, E') (-f) := diff --git a/src/group_theory/order_of_element.lean b/src/group_theory/order_of_element.lean index 7bf7c11411d5a..13773d1c59379 100644 --- a/src/group_theory/order_of_element.lean +++ b/src/group_theory/order_of_element.lean @@ -314,9 +314,10 @@ order_of_pos_iff.mp $ pos_of_dvd_of_pos h.order_of_mul_dvd_mul_order_of $ mul_pos (order_of_pos' hx) (order_of_pos' hy) /-- If each prime factor of `order_of x` has higher multiplicity in `order_of y`, and `x` commutes - with `y`, then the order of `x * y` is the same as the order of `y`. -/ -@[to_additive "If each prime factor of `add_order_of x` has higher multiplicity in `add_order_of y`, -and `x` commutes with `y`, then the order of `x * y` is the same as the order of `y`."] + with `y`, then `x * y` has the same order as `y`. -/ +@[to_additive add_order_of_add_eq_right_of_forall_prime_mul_dvd "If each prime factor of + `add_order_of x` has higher multiplicity in `add_order_of y`, and `x` commutes with `y`, + then `x + y` has the same order as `y`."] lemma order_of_mul_eq_right_of_forall_prime_mul_dvd (hy : is_of_fin_order y) (hdvd : ∀ p : ℕ, p.prime → p ∣ order_of x → (p * order_of x) ∣ order_of y) : @@ -328,7 +329,7 @@ begin { exact trans h.order_of_mul_dvd_lcm (lcm_dvd hxy dvd_rfl) }, refine λ p hp hpy hd, hp.ne_one _, rw [← nat.dvd_one, ← mul_dvd_mul_iff_right hoy.ne', one_mul, ← dvd_div_iff hpy], - refine trans (order_of_dvd_lcm_mul h) (lcm_dvd_iff.2 ⟨(dvd_div_iff hpy).2 _, hd⟩), + refine trans (order_of_dvd_lcm_mul h) (lcm_dvd ((dvd_div_iff hpy).2 _) hd), by_cases p ∣ order_of x, exacts [hdvd p hp h, (hp.coprime_iff_not_dvd.2 h).mul_dvd_of_dvd_of_dvd hpy hxy], end diff --git a/src/tactic/print_sorry.lean b/src/tactic/print_sorry.lean index f380ddec871fb..d2e6421a396a9 100644 --- a/src/tactic/print_sorry.lean +++ b/src/tactic/print_sorry.lean @@ -17,7 +17,7 @@ Other searches through the environment can be done using `tactic.find_all_exprs` namespace tactic -/-- Auxilliary data type for `tactic.find_all_exprs` -/ +/-- Auxiliary data type for `tactic.find_all_exprs` -/ meta structure find_all_expr_data := (matching_subexpr : bool) -- this declaration contains a subexpression on which the test passes (test_passed : bool) -- the search has found a matching subexpression somewhere @@ -26,7 +26,7 @@ meta structure find_all_expr_data := (name_map : name_map bool) -- all data (direct_descendants : name_set) -- direct descendants of a declaration -/-- Auxilliary declaration for `tactic.find_all_exprs`. +/-- Auxiliary declaration for `tactic.find_all_exprs`. Traverse all declarations occurring in the declaration with the given name, excluding declarations `n` such that `g n` is true (and all their descendants), From 3813d4ea1c6a34dbb472de66e73b8c6855b03964 Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Fri, 23 Dec 2022 11:28:42 +0000 Subject: [PATCH 0102/1291] chore(data/nat/modeq): split out lemmas about lists (#18004) By splitting out some lemmas relating list-rotation and modular arithmetic, the files `data.nat.modeq` and `data.int.modeq` become portable now. --- archive/miu_language/decision_nec.lean | 1 + src/data/list/modeq.lean | 68 ++++++++++++++++++++++++++ src/data/nat/modeq.lean | 60 ----------------------- src/group_theory/perm/cycle/type.lean | 1 + 4 files changed, 70 insertions(+), 60 deletions(-) create mode 100644 src/data/list/modeq.lean diff --git a/archive/miu_language/decision_nec.lean b/archive/miu_language/decision_nec.lean index 8c540078509f6..d70116a20b054 100644 --- a/archive/miu_language/decision_nec.lean +++ b/archive/miu_language/decision_nec.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Gihan Marasingha -/ import .basic +import data.list.count import data.nat.modeq import tactic.ring diff --git a/src/data/list/modeq.lean b/src/data/list/modeq.lean new file mode 100644 index 0000000000000..10121abd06831 --- /dev/null +++ b/src/data/list/modeq.lean @@ -0,0 +1,68 @@ +/- +Copyright (c) 2018 Chris Hughes. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Chris Hughes +-/ +import data.nat.modeq +import data.list.rotate + +/-! # List rotation and modular arithmetic -/ + +namespace list +variable {α : Type*} + +lemma nth_rotate : ∀ {l : list α} {n m : ℕ} (hml : m < l.length), + (l.rotate n).nth m = l.nth ((m + n) % l.length) +| [] n m hml := (nat.not_lt_zero _ hml).elim +| l 0 m hml := by simp [nat.mod_eq_of_lt hml] +| (a::l) (n+1) m hml := +have h₃ : m < list.length (l ++ [a]), by simpa using hml, +(lt_or_eq_of_le (nat.le_of_lt_succ $ nat.mod_lt (m + n) + (lt_of_le_of_lt (nat.zero_le _) hml))).elim +(λ hml', + have h₁ : (m + (n + 1)) % ((a :: l : list α).length) = + (m + n) % ((a :: l : list α).length) + 1, + from calc (m + (n + 1)) % (l.length + 1) = + ((m + n) % (l.length + 1) + 1) % (l.length + 1) : + add_assoc m n 1 ▸ nat.modeq.add_right 1 (nat.mod_mod _ _).symm + ... = (m + n) % (l.length + 1) + 1 : nat.mod_eq_of_lt (nat.succ_lt_succ hml'), + have h₂ : (m + n) % (l ++ [a]).length < l.length, by simpa [nat.add_one] using hml', + by rw [list.rotate_cons_succ, nth_rotate h₃, list.nth_append h₂, h₁, list.nth]; simp) +(λ hml', + have h₁ : (m + (n + 1)) % (l.length + 1) = 0, + from calc (m + (n + 1)) % (l.length + 1) = (l.length + 1) % (l.length + 1) : + add_assoc m n 1 ▸ nat.modeq.add_right 1 + (hml'.trans (nat.mod_eq_of_lt (nat.lt_succ_self _)).symm) + ... = 0 : by simp, + by rw [list.length, list.rotate_cons_succ, nth_rotate h₃, list.length_append, + list.length_cons, list.length, zero_add, hml', h₁, list.nth_concat_length]; refl) + +lemma rotate_eq_self_iff_eq_repeat [hα : nonempty α] : ∀ {l : list α}, + (∀ n, l.rotate n = l) ↔ ∃ a, l = list.repeat a l.length +| [] := ⟨λ h, nonempty.elim hα (λ a, ⟨a, by simp⟩), by simp⟩ +| (a::l) := +⟨λ h, ⟨a, list.ext_le (by simp) $ λ n hn h₁, + begin + rw [← option.some_inj, ← list.nth_le_nth], + conv {to_lhs, rw ← h ((list.length (a :: l)) - n)}, + rw [nth_rotate hn, add_tsub_cancel_of_le (le_of_lt hn), + nat.mod_self, nth_le_repeat], refl + end⟩, + λ ⟨a, ha⟩ n, ha.symm ▸ list.ext_le (by simp) + (λ m hm h, + have hm' : (m + n) % (list.repeat a (list.length (a :: l))).length < list.length (a :: l), + by rw list.length_repeat; exact nat.mod_lt _ (nat.succ_pos _), + by rw [nth_le_repeat, ← option.some_inj, ← list.nth_le_nth, nth_rotate h, list.nth_le_nth, + nth_le_repeat]; simp * at *)⟩ + +lemma rotate_repeat (a : α) (n : ℕ) (k : ℕ) : + (list.repeat a n).rotate k = list.repeat a n := +let h : nonempty α := ⟨a⟩ in by exactI rotate_eq_self_iff_eq_repeat.mpr ⟨a, by rw length_repeat⟩ k + +lemma rotate_one_eq_self_iff_eq_repeat [nonempty α] {l : list α} : + l.rotate 1 = l ↔ ∃ a : α, l = list.repeat a l.length := +⟨λ h, rotate_eq_self_iff_eq_repeat.mp (λ n, nat.rec l.rotate_zero + (λ n hn, by rwa [nat.succ_eq_add_one, ←l.rotate_rotate, hn]) n), + λ h, rotate_eq_self_iff_eq_repeat.mpr h 1⟩ + +end list diff --git a/src/data/nat/modeq.lean b/src/data/nat/modeq.lean index 59979347e8aaa..1f1dfc5bf2e15 100644 --- a/src/data/nat/modeq.lean +++ b/src/data/nat/modeq.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ import data.int.gcd -import data.list.rotate import tactic.abel /-! @@ -425,62 +424,3 @@ have help : ∀ (m : ℕ), m < 4 → m % 2 = 1 → m = 1 ∨ m = 3 := dec_trivia λ h, or.dcases_on h odd_of_mod_four_eq_one odd_of_mod_four_eq_three⟩ end nat - -namespace list -variable {α : Type*} - -lemma nth_rotate : ∀ {l : list α} {n m : ℕ} (hml : m < l.length), - (l.rotate n).nth m = l.nth ((m + n) % l.length) -| [] n m hml := (nat.not_lt_zero _ hml).elim -| l 0 m hml := by simp [nat.mod_eq_of_lt hml] -| (a::l) (n+1) m hml := -have h₃ : m < list.length (l ++ [a]), by simpa using hml, -(lt_or_eq_of_le (nat.le_of_lt_succ $ nat.mod_lt (m + n) - (lt_of_le_of_lt (nat.zero_le _) hml))).elim -(λ hml', - have h₁ : (m + (n + 1)) % ((a :: l : list α).length) = - (m + n) % ((a :: l : list α).length) + 1, - from calc (m + (n + 1)) % (l.length + 1) = - ((m + n) % (l.length + 1) + 1) % (l.length + 1) : - add_assoc m n 1 ▸ nat.modeq.add_right 1 (nat.mod_mod _ _).symm - ... = (m + n) % (l.length + 1) + 1 : nat.mod_eq_of_lt (nat.succ_lt_succ hml'), - have h₂ : (m + n) % (l ++ [a]).length < l.length, by simpa [nat.add_one] using hml', - by rw [list.rotate_cons_succ, nth_rotate h₃, list.nth_append h₂, h₁, list.nth]; simp) -(λ hml', - have h₁ : (m + (n + 1)) % (l.length + 1) = 0, - from calc (m + (n + 1)) % (l.length + 1) = (l.length + 1) % (l.length + 1) : - add_assoc m n 1 ▸ nat.modeq.add_right 1 - (hml'.trans (nat.mod_eq_of_lt (nat.lt_succ_self _)).symm) - ... = 0 : by simp, - by rw [list.length, list.rotate_cons_succ, nth_rotate h₃, list.length_append, - list.length_cons, list.length, zero_add, hml', h₁, list.nth_concat_length]; refl) - -lemma rotate_eq_self_iff_eq_repeat [hα : nonempty α] : ∀ {l : list α}, - (∀ n, l.rotate n = l) ↔ ∃ a, l = list.repeat a l.length -| [] := ⟨λ h, nonempty.elim hα (λ a, ⟨a, by simp⟩), by simp⟩ -| (a::l) := -⟨λ h, ⟨a, list.ext_le (by simp) $ λ n hn h₁, - begin - rw [← option.some_inj, ← list.nth_le_nth], - conv {to_lhs, rw ← h ((list.length (a :: l)) - n)}, - rw [nth_rotate hn, add_tsub_cancel_of_le (le_of_lt hn), - nat.mod_self, nth_le_repeat], refl - end⟩, - λ ⟨a, ha⟩ n, ha.symm ▸ list.ext_le (by simp) - (λ m hm h, - have hm' : (m + n) % (list.repeat a (list.length (a :: l))).length < list.length (a :: l), - by rw list.length_repeat; exact nat.mod_lt _ (nat.succ_pos _), - by rw [nth_le_repeat, ← option.some_inj, ← list.nth_le_nth, nth_rotate h, list.nth_le_nth, - nth_le_repeat]; simp * at *)⟩ - -lemma rotate_repeat (a : α) (n : ℕ) (k : ℕ) : - (list.repeat a n).rotate k = list.repeat a n := -let h : nonempty α := ⟨a⟩ in by exactI rotate_eq_self_iff_eq_repeat.mpr ⟨a, by rw length_repeat⟩ k - -lemma rotate_one_eq_self_iff_eq_repeat [nonempty α] {l : list α} : - l.rotate 1 = l ↔ ∃ a : α, l = list.repeat a l.length := -⟨λ h, rotate_eq_self_iff_eq_repeat.mp (λ n, nat.rec l.rotate_zero - (λ n hn, by rwa [nat.succ_eq_add_one, ←l.rotate_rotate, hn]) n), - λ h, rotate_eq_self_iff_eq_repeat.mpr h 1⟩ - -end list diff --git a/src/group_theory/perm/cycle/type.lean b/src/group_theory/perm/cycle/type.lean index afc73b7b01826..52bb7c485ca3d 100644 --- a/src/group_theory/perm/cycle/type.lean +++ b/src/group_theory/perm/cycle/type.lean @@ -6,6 +6,7 @@ Authors: Thomas Browning import algebra.gcd_monoid.multiset import combinatorics.partition +import data.list.modeq import group_theory.perm.cycle.basic import ring_theory.int.basic import tactic.linarith From ac1e2e08841216b84cd59a7b601ee87f8c916b86 Mon Sep 17 00:00:00 2001 From: Anatole Dedecker Date: Fri, 23 Dec 2022 16:07:39 +0000 Subject: [PATCH 0103/1291] feat(analysis/locally_convex/bounded): characterize boundedness using sequences (#17688) --- src/analysis/locally_convex/bounded.lean | 62 +++++++++++++++++++++++- src/order/filter/basic.lean | 12 +++++ 2 files changed, 73 insertions(+), 1 deletion(-) diff --git a/src/analysis/locally_convex/bounded.lean b/src/analysis/locally_convex/bounded.lean index d26f8bcfcdcaa..a2b3d5c7767ba 100644 --- a/src/analysis/locally_convex/bounded.lean +++ b/src/analysis/locally_convex/bounded.lean @@ -26,6 +26,11 @@ absorbs `s`. * `bornology.is_vonN_bounded.of_topological_space_le`: A coarser topology admits more von Neumann-bounded sets. * `bornology.is_vonN_bounded.image`: A continuous linear image of a bounded set is bounded. +* `bornology.is_vonN_bounded_iff_smul_tendsto_zero`: Given any sequence `ε` of scalars which tends + to `𝓝[≠] 0`, we have that a set `S` is bounded if and only if for any sequence `x : ℕ → S`, + `ε • x` tends to 0. This shows that bounded sets are completely determined by sequences, which is + the key fact for proving that sequential continuity implies continuity for linear maps defined on + a bornological space ## References @@ -35,7 +40,7 @@ von Neumann-bounded sets. variables {𝕜 𝕜' E E' F ι : Type*} -open filter +open set filter open_locale topological_space pointwise namespace bornology @@ -128,6 +133,61 @@ end end image +section sequence + +variables {𝕝 : Type*} [normed_field 𝕜] [nontrivially_normed_field 𝕝] [add_comm_group E] [module 𝕜 E] + [module 𝕝 E] [topological_space E] [has_continuous_smul 𝕝 E] + +lemma is_vonN_bounded.smul_tendsto_zero {S : set E} {ε : ι → 𝕜} {x : ι → E} {l : filter ι} + (hS : is_vonN_bounded 𝕜 S) (hxS : ∀ᶠ n in l, x n ∈ S) (hε : tendsto ε l (𝓝 0)) : + tendsto (ε • x) l (𝓝 0) := +begin + rw tendsto_def at *, + intros V hV, + rcases hS hV with ⟨r, r_pos, hrS⟩, + filter_upwards [hxS, hε _ (metric.ball_mem_nhds 0 $ inv_pos.mpr r_pos)] with n hnS hnr, + by_cases this : ε n = 0, + { simp [this, mem_of_mem_nhds hV] }, + { rw [mem_preimage, mem_ball_zero_iff, lt_inv (norm_pos_iff.mpr this) r_pos, ← norm_inv] at hnr, + rw [mem_preimage, pi.smul_apply', ← set.mem_inv_smul_set_iff₀ this], + exact hrS _ (hnr.le) hnS }, +end + +lemma is_vonN_bounded_of_smul_tendsto_zero {ε : ι → 𝕝} {l : filter ι} [l.ne_bot] + (hε : ∀ᶠ n in l, ε n ≠ 0) {S : set E} + (H : ∀ x : ι → E, (∀ n, x n ∈ S) → tendsto (ε • x) l (𝓝 0)) : + is_vonN_bounded 𝕝 S := +begin + rw (nhds_basis_balanced 𝕝 E).is_vonN_bounded_basis_iff, + by_contra' H', + rcases H' with ⟨V, ⟨hV, hVb⟩, hVS⟩, + have : ∀ᶠ n in l, ∃ x : S, (ε n) • (x : E) ∉ V, + { filter_upwards [hε] with n hn, + rw absorbs at hVS, + push_neg at hVS, + rcases hVS _ (norm_pos_iff.mpr $ inv_ne_zero hn) with ⟨a, haε, haS⟩, + rcases set.not_subset.mp haS with ⟨x, hxS, hx⟩, + refine ⟨⟨x, hxS⟩, λ hnx, _⟩, + rw ← set.mem_inv_smul_set_iff₀ hn at hnx, + exact hx (hVb.smul_mono haε hnx) }, + rcases this.choice with ⟨x, hx⟩, + refine filter.frequently_false l (filter.eventually.frequently _), + filter_upwards [hx, (H (coe ∘ x) (λ n, (x n).2)).eventually (eventually_mem_set.mpr hV)] + using λ n, id +end + +/-- Given any sequence `ε` of scalars which tends to `𝓝[≠] 0`, we have that a set `S` is bounded + if and only if for any sequence `x : ℕ → S`, `ε • x` tends to 0. This actually works for any + indexing type `ι`, but in the special case `ι = ℕ` we get the important fact that convergent + sequences fully characterize bounded sets. -/ +lemma is_vonN_bounded_iff_smul_tendsto_zero {ε : ι → 𝕝} {l : filter ι} [l.ne_bot] + (hε : tendsto ε l (𝓝[≠] 0)) {S : set E} : + is_vonN_bounded 𝕝 S ↔ ∀ x : ι → E, (∀ n, x n ∈ S) → tendsto (ε • x) l (𝓝 0) := +⟨λ hS x hxS, hS.smul_tendsto_zero (eventually_of_forall hxS) (le_trans hε nhds_within_le_nhds), + is_vonN_bounded_of_smul_tendsto_zero (hε self_mem_nhds_within)⟩ + +end sequence + section normed_field variables [normed_field 𝕜] [add_comm_group E] [module 𝕜 E] diff --git a/src/order/filter/basic.lean b/src/order/filter/basic.lean index 125aef03200b7..79538fcb2be36 100644 --- a/src/order/filter/basic.lean +++ b/src/order/filter/basic.lean @@ -1232,6 +1232,18 @@ lemma frequently_supr {p : α → Prop} {fs : β → filter α} : (∃ᶠ x in (⨆ b, fs b), p x) ↔ (∃ b, ∃ᶠ x in fs b, p x) := by simp [filter.frequently, -not_eventually, not_forall] +lemma eventually.choice {r : α → β → Prop} {l : filter α} + [l.ne_bot] (h : ∀ᶠ x in l, ∃ y, r x y) : ∃ f : α → β, ∀ᶠ x in l, r x (f x) := +begin + classical, + use (λ x, if hx : ∃ y, r x y then classical.some hx + else classical.some (classical.some_spec h.exists)), + filter_upwards [h], + intros x hx, + rw dif_pos hx, + exact classical.some_spec hx +end + /-! ### Relation “eventually equal” -/ From 717c13217a49955a2955133f5b3991b0a9669878 Mon Sep 17 00:00:00 2001 From: Gabriel Ebner Date: Fri, 23 Dec 2022 22:40:21 +0000 Subject: [PATCH 0104/1291] feat(*): bump to lean 3.50.2 (#18007) --- leanpkg.toml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/leanpkg.toml b/leanpkg.toml index c50815f219262..ac4a167bb424a 100644 --- a/leanpkg.toml +++ b/leanpkg.toml @@ -1,7 +1,7 @@ [package] name = "mathlib" version = "0.1" -lean_version = "leanprover-community/lean:3.50.1" +lean_version = "leanprover-community/lean:3.50.2" path = "src" [dependencies] From 09258fb7f75d741b7eda9fa18d5c869e2135d9f1 Mon Sep 17 00:00:00 2001 From: datokrat Date: Sat, 24 Dec 2022 02:09:48 +0000 Subject: [PATCH 0105/1291] feat(analysis/normed_space/affine_isometry, linear_algebra/affine_space/affine_equiv): restrict affine isometry to isometry equivalence (#17522) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The main result in this commit is `affine_subspace.isometry_equiv_map`: Given an affine isometry, each of its affine subspaces is affine isometry equivalent to its image. `isometry_equiv_map` returns this isometry equivalence. The two other most significant results that are proved on the way are: - `affine_subspace.equiv_map_of_injective`: Restricts an injective affine map to an affine equivalence of a subspace to its image (used by `isometry_equiv_map`) - `affine_equiv.of_bijective`: obtain an affine equivalence from a bijective affine map The construction uses the new definition `affine_equiv.of_bijective` that makes an affine equivalence of a bijective affine map. Co-authored-by: Eric Wieser Co-authored-by: Yaël Dillies Co-authored-by: Eric Wieser --- src/analysis/convex/side.lean | 4 +- .../normed_space/add_torsor_bases.lean | 2 +- .../normed_space/affine_isometry.lean | 52 +++++++++++++++++++ src/analysis/normed_space/banach.lean | 2 +- .../affine_space/affine_equiv.lean | 10 ++++ .../affine_space/affine_map.lean | 8 ++- .../affine_space/independent.lean | 2 +- src/linear_algebra/affine_space/restrict.lean | 6 +++ 8 files changed, 79 insertions(+), 7 deletions(-) diff --git a/src/analysis/convex/side.lean b/src/analysis/convex/side.lean index 3b2f0899197eb..827cd0dff4c6b 100644 --- a/src/analysis/convex/side.lean +++ b/src/analysis/convex/side.lean @@ -75,7 +75,7 @@ begin rcases hfp₁ with ⟨p₁, hp₁, rfl⟩, rcases hfp₂ with ⟨p₂, hp₂, rfl⟩, refine ⟨p₁, hp₁, p₂, hp₂, _⟩, - simp_rw [←linear_map_vsub, (f.injective_iff_linear_injective.2 hf).same_ray_map_iff] at h, + simp_rw [←linear_map_vsub, (f.linear_injective_iff.2 hf).same_ray_map_iff] at h, exact h end @@ -111,7 +111,7 @@ begin rcases hfp₁ with ⟨p₁, hp₁, rfl⟩, rcases hfp₂ with ⟨p₂, hp₂, rfl⟩, refine ⟨p₁, hp₁, p₂, hp₂, _⟩, - simp_rw [←linear_map_vsub, (f.injective_iff_linear_injective.2 hf).same_ray_map_iff] at h, + simp_rw [←linear_map_vsub, (f.linear_injective_iff.2 hf).same_ray_map_iff] at h, exact h end diff --git a/src/analysis/normed_space/add_torsor_bases.lean b/src/analysis/normed_space/add_torsor_bases.lean index e41165992a954..e5b9b25a7e6a0 100644 --- a/src/analysis/normed_space/add_torsor_bases.lean +++ b/src/analysis/normed_space/add_torsor_bases.lean @@ -33,7 +33,7 @@ include E lemma is_open_map_barycentric_coord [nontrivial ι] (b : affine_basis ι 𝕜 P) (i : ι) : is_open_map (b.coord i) := affine_map.is_open_map_linear_iff.mp $ (b.coord i).linear.is_open_map_of_finite_dimensional $ - (b.coord i).surjective_iff_linear_surjective.mpr (b.surjective_coord i) + (b.coord i).linear_surjective_iff.mpr (b.surjective_coord i) variables [finite_dimensional 𝕜 E] (b : affine_basis ι 𝕜 P) diff --git a/src/analysis/normed_space/affine_isometry.lean b/src/analysis/normed_space/affine_isometry.lean index d95debabb91d5..34389bb722cb6 100644 --- a/src/analysis/normed_space/affine_isometry.lean +++ b/src/analysis/normed_space/affine_isometry.lean @@ -6,6 +6,7 @@ Authors: Heather Macbeth import analysis.normed_space.linear_isometry import analysis.normed.group.add_torsor import analysis.normed_space.basic +import linear_algebra.affine_space.restrict /-! # Affine isometries @@ -587,3 +588,54 @@ begin rw this, simp only [homeomorph.comp_is_open_map_iff, homeomorph.comp_is_open_map_iff'], end + +local attribute [instance, nolint fails_quickly] affine_subspace.nonempty_map + +include V₁ +omit V + +namespace affine_subspace + +/-- +An affine subspace is isomorphic to its image under an injective affine map. +This is the affine version of `submodule.equiv_map_of_injective`. +-/ +@[simps] +noncomputable def equiv_map_of_injective (E: affine_subspace 𝕜 P₁) [nonempty E] + (φ : P₁ →ᵃ[𝕜] P₂) (hφ : function.injective φ) : E ≃ᵃ[𝕜] E.map φ := +{ linear := + (E.direction.equiv_map_of_injective φ.linear (φ.linear_injective_iff.mpr hφ)).trans + (linear_equiv.of_eq _ _ (affine_subspace.map_direction _ _).symm), + map_vadd' := λ p v, subtype.ext $ φ.map_vadd p v, + .. equiv.set.image _ (E : set P₁) hφ } + +/-- +Restricts an affine isometry to an affine isometry equivalence between a nonempty affine +subspace `E` and its image. + +This is an isometry version of `affine_subspace.equiv_map`, having a stronger premise and a stronger +conclusion. +-/ +noncomputable def isometry_equiv_map + (φ : P₁ →ᵃⁱ[𝕜] P₂) (E : affine_subspace 𝕜 P₁) [nonempty E] : E ≃ᵃⁱ[𝕜] E.map φ.to_affine_map := +⟨E.equiv_map_of_injective φ.to_affine_map φ.injective, (λ _, φ.norm_map _)⟩ + +@[simp] +lemma isometry_equiv_map.apply_symm_apply + {E : affine_subspace 𝕜 P₁} [nonempty E] + {φ : P₁ →ᵃⁱ[𝕜] P₂} (x : E.map φ.to_affine_map) : + φ ((E.isometry_equiv_map φ).symm x) = x := +congr_arg coe $ (E.isometry_equiv_map φ).apply_symm_apply _ + +@[simp] +lemma isometry_equiv_map.coe_apply + (φ : P₁ →ᵃⁱ[𝕜] P₂) (E : affine_subspace 𝕜 P₁) [nonempty E] (g: E) : + ↑(E.isometry_equiv_map φ g) = φ g := rfl + +@[simp] +lemma isometry_equiv_map.to_affine_map_eq + (φ : P₁ →ᵃⁱ[𝕜] P₂) (E : affine_subspace 𝕜 P₁) [nonempty E] : + (E.isometry_equiv_map φ).to_affine_map = E.equiv_map_of_injective φ.to_affine_map φ.injective := +rfl + +end affine_subspace diff --git a/src/analysis/normed_space/banach.lean b/src/analysis/normed_space/banach.lean index e34f631b63f55..c900d45d51f63 100644 --- a/src/analysis/normed_space/banach.lean +++ b/src/analysis/normed_space/banach.lean @@ -250,7 +250,7 @@ lemma _root_.affine_map.is_open_map {P Q : Type*} is_open_map f := affine_map.is_open_map_linear_iff.mp $ continuous_linear_map.is_open_map { cont := affine_map.continuous_linear_iff.mpr hf, .. f.linear } - (f.surjective_iff_linear_surjective.mpr surj) + (f.linear_surjective_iff.mpr surj) /-! ### Applications of the Banach open mapping theorem -/ diff --git a/src/linear_algebra/affine_space/affine_equiv.lean b/src/linear_algebra/affine_space/affine_equiv.lean index 0f740a267e3e8..fb25afd2fccdc 100644 --- a/src/linear_algebra/affine_space/affine_equiv.lean +++ b/src/linear_algebra/affine_space/affine_equiv.lean @@ -160,6 +160,16 @@ protected lemma bijective (e : P₁ ≃ᵃ[k] P₂) : bijective e := e.to_equiv. protected lemma surjective (e : P₁ ≃ᵃ[k] P₂) : surjective e := e.to_equiv.surjective protected lemma injective (e : P₁ ≃ᵃ[k] P₂) : injective e := e.to_equiv.injective +/-- Bijective affine maps are affine isomorphisms. -/ +@[simps] +noncomputable def of_bijective {φ : P₁ →ᵃ[k] P₂} (hφ : function.bijective φ) : P₁ ≃ᵃ[k] P₂ := +{ linear := linear_equiv.of_bijective φ.linear (φ.linear_bijective_iff.mpr hφ), + map_vadd' := φ.map_vadd, + ..(equiv.of_bijective _ hφ) } + +lemma of_bijective.symm_eq {φ : P₁ →ᵃ[k] P₂} (hφ : function.bijective φ) : + (of_bijective hφ).symm.to_equiv = (equiv.of_bijective _ hφ).symm := rfl + @[simp] lemma range_eq (e : P₁ ≃ᵃ[k] P₂) : range e = univ := e.surjective.range_eq @[simp] lemma apply_symm_apply (e : P₁ ≃ᵃ[k] P₂) (p : P₂) : e (e.symm p) = p := diff --git a/src/linear_algebra/affine_space/affine_map.lean b/src/linear_algebra/affine_space/affine_map.lean index 15b92296a23b5..f6817af397589 100644 --- a/src/linear_algebra/affine_space/affine_map.lean +++ b/src/linear_algebra/affine_space/affine_map.lean @@ -328,7 +328,7 @@ instance : monoid (P1 →ᵃ[k] P1) := include V2 -@[simp] lemma injective_iff_linear_injective (f : P1 →ᵃ[k] P2) : +@[simp] lemma linear_injective_iff (f : P1 →ᵃ[k] P2) : function.injective f.linear ↔ function.injective f := begin obtain ⟨p⟩ := (infer_instance : nonempty P1), @@ -337,7 +337,7 @@ begin rw [h, equiv.comp_injective, equiv.injective_comp], end -@[simp] lemma surjective_iff_linear_surjective (f : P1 →ᵃ[k] P2) : +@[simp] lemma linear_surjective_iff (f : P1 →ᵃ[k] P2) : function.surjective f.linear ↔ function.surjective f := begin obtain ⟨p⟩ := (infer_instance : nonempty P1), @@ -346,6 +346,10 @@ begin rw [h, equiv.comp_surjective, equiv.surjective_comp], end +@[simp] lemma linear_bijective_iff (f : P1 →ᵃ[k] P2) : + function.bijective f.linear ↔ function.bijective f := +and_congr f.linear_injective_iff f.linear_surjective_iff + lemma image_vsub_image {s t : set P1} (f : P1 →ᵃ[k] P2) : (f '' s) -ᵥ (f '' t) = f.linear '' (s -ᵥ t) := begin diff --git a/src/linear_algebra/affine_space/independent.lean b/src/linear_algebra/affine_space/independent.lean index 2e485e13f0703..1bd97d4143948 100644 --- a/src/linear_algebra/affine_space/independent.lean +++ b/src/linear_algebra/affine_space/independent.lean @@ -365,7 +365,7 @@ begin rw affine_independent_iff_linear_independent_vsub k p i at hai, simp_rw [affine_independent_iff_linear_independent_vsub k (f ∘ p) i, function.comp_app, ← f.linear_map_vsub], - have hf' : f.linear.ker = ⊥, { rwa [linear_map.ker_eq_bot, f.injective_iff_linear_injective], }, + have hf' : f.linear.ker = ⊥, { rwa [linear_map.ker_eq_bot, f.linear_injective_iff], }, exact linear_independent.map' hai f.linear hf', end diff --git a/src/linear_algebra/affine_space/restrict.lean b/src/linear_algebra/affine_space/restrict.lean index 029252a1b42d3..5a82ba53e4136 100644 --- a/src/linear_algebra/affine_space/restrict.lean +++ b/src/linear_algebra/affine_space/restrict.lean @@ -99,3 +99,9 @@ begin obtain ⟨y, hy, rfl⟩ := hx, exact ⟨⟨y, hy⟩, rfl⟩, end + +lemma affine_map.restrict.bijective + {E : affine_subspace k P₁} [nonempty E] + {φ : P₁ →ᵃ[k] P₂} (hφ : function.injective φ) : + function.bijective (φ.restrict (le_refl (E.map φ))) := +⟨affine_map.restrict.injective hφ _, affine_map.restrict.surjective _ rfl⟩ From 930133160e24036d5242039fe4972407cd4f1222 Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Sat, 24 Dec 2022 18:13:43 +0000 Subject: [PATCH 0106/1291] feat(algebra/cubic_discriminant): extract useful lemma from eq_sum_three_roots (#17999) --- src/algebra/cubic_discriminant.lean | 14 +++++++++++--- 1 file changed, 11 insertions(+), 3 deletions(-) diff --git a/src/algebra/cubic_discriminant.lean b/src/algebra/cubic_discriminant.lean index 2ac5858261166..e10263ae6e55c 100644 --- a/src/algebra/cubic_discriminant.lean +++ b/src/algebra/cubic_discriminant.lean @@ -55,6 +55,16 @@ variables {P Q : cubic R} {a b c d a' b' c' d' : R} [semiring R] /-- Convert a cubic polynomial to a polynomial. -/ def to_poly (P : cubic R) : R[X] := C P.a * X ^ 3 + C P.b * X ^ 2 + C P.c * X + C P.d +theorem C_mul_prod_X_sub_C_eq [comm_ring S] {w x y z : S} : + C w * (X - C x) * (X - C y) * (X - C z) + = to_poly ⟨w, w * -(x + y + z), w * (x * y + x * z + y * z), w * -(x * y * z)⟩ := +by { simp only [to_poly, C_neg, C_add, C_mul], ring1 } + +theorem prod_X_sub_C_eq [comm_ring S] {x y z : S} : + (X - C x) * (X - C y) * (X - C z) + = to_poly ⟨1, -(x + y + z), (x * y + x * z + y * z), -(x * y * z)⟩ := +by rw [← one_mul $ X - C x, ← C_1, C_mul_prod_X_sub_C_eq, one_mul, one_mul, one_mul] + /-! ### Coefficients -/ section coeff @@ -382,9 +392,7 @@ theorem eq_sum_three_roots (ha : P.a ≠ 0) (h3 : (map φ P).roots = {x, y, z}) begin apply_fun to_poly, any_goals { exact λ P Q, (to_poly_injective P Q).mp }, - rw [eq_prod_three_roots ha h3, to_poly], - simp only [C_neg, C_add, C_mul], - ring1 + rw [eq_prod_three_roots ha h3, C_mul_prod_X_sub_C_eq] end theorem b_eq_three_roots (ha : P.a ≠ 0) (h3 : (map φ P).roots = {x, y, z}) : From 8acf2263798568a24cc02ad7a22673b359b42bc3 Mon Sep 17 00:00:00 2001 From: Niels Voss Date: Sat, 24 Dec 2022 19:38:33 +0000 Subject: [PATCH 0107/1291] feat(algebra/geom_sum): divisibility of sums and differences of powers (#17996) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit If `n : ℕ` and `x` and `y` belong to a `comm_ring`, then `x - y ∣ x ^ n - y ^ n`. If `n` is odd, then `x + y ∣ x ^ n + y ^ n`. These results follow from `geom_sum₂_mul`. There are additional theorems for when `x` and `y` are natural numbers. Co-authored-by: Niels Voss --- src/algebra/geom_sum.lean | 23 +++++++++++++++++++++++ 1 file changed, 23 insertions(+) diff --git a/src/algebra/geom_sum.lean b/src/algebra/geom_sum.lean index d129841e1be0e..2978020084607 100644 --- a/src/algebra/geom_sum.lean +++ b/src/algebra/geom_sum.lean @@ -172,6 +172,29 @@ theorem geom_sum₂_mul [comm_ring α] (x y : α) (n : ℕ) : (∑ i in range n, x ^ i * (y ^ (n - 1 - i))) * (x - y) = x ^ n - y ^ n := (commute.all x y).geom_sum₂_mul n +theorem sub_dvd_pow_sub_pow [comm_ring α] (x y : α) (n : ℕ) : x - y ∣ x ^ n - y ^ n := + dvd.intro_left _ (geom_sum₂_mul x y n) + +theorem nat_sub_dvd_pow_sub_pow (x y n : ℕ) : x - y ∣ x ^ n - y ^ n := +begin + cases le_or_lt y x with h, + { have : y ^ n ≤ x ^ n := nat.pow_le_pow_of_le_left h _, + exact_mod_cast sub_dvd_pow_sub_pow (x : ℤ) ↑y n }, + { have : x ^ n ≤ y ^ n := nat.pow_le_pow_of_le_left h.le _, + exact (nat.sub_eq_zero_of_le this).symm ▸ dvd_zero (x - y) } +end + +theorem odd.add_dvd_pow_add_pow [comm_ring α] (x y : α) {n : ℕ} (h : odd n) : + x + y ∣ x ^ n + y ^ n := +begin + have h₁ := geom_sum₂_mul x (-y) n, + rw [odd.neg_pow h y, sub_neg_eq_add, sub_neg_eq_add] at h₁, + exact dvd.intro_left _ h₁, +end + +theorem odd.nat_add_dvd_pow_add_pow (x y : ℕ) {n : ℕ} (h : odd n) : x + y ∣ x ^ n + y ^ n := +by exact_mod_cast odd.add_dvd_pow_add_pow (x : ℤ) ↑y h + theorem geom_sum_mul [ring α] (x : α) (n : ℕ) : (∑ i in range n, x ^ i) * (x - 1) = x ^ n - 1 := begin From e9be2fa75faa22892937c275e27a91cd558cf8c0 Mon Sep 17 00:00:00 2001 From: loefflerd Date: Mon, 26 Dec 2022 08:07:58 +0000 Subject: [PATCH 0108/1291] =?UTF-8?q?feat(analysis/complex):=20integration?= =?UTF-8?q?=20and=20differentiation=20`=E2=84=9D=20=E2=86=92=20=E2=84=82`?= =?UTF-8?q?=20(#17989)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit In Fourier theory it's often necessary to work with integrals and derivatives of functions `ℝ → ℂ`. This adds a couple of shortcuts for doing so, and uses them to slightly simplify some code already in the library. (Cherry-picked out of my Riemann zeta values project.) --- src/analysis/complex/real_deriv.lean | 18 ++++++++++++++-- src/analysis/special_functions/gamma.lean | 21 ++++++++----------- src/analysis/special_functions/integrals.lean | 12 +++++------ .../integral/interval_integral.lean | 4 ++++ 4 files changed, 34 insertions(+), 21 deletions(-) diff --git a/src/analysis/complex/real_deriv.lean b/src/analysis/complex/real_deriv.lean index c06c96100c8cc..f3cf1837f66c2 100644 --- a/src/analysis/complex/real_deriv.lean +++ b/src/analysis/complex/real_deriv.lean @@ -51,8 +51,9 @@ begin simpa using (C.comp z (B.comp z A)).has_strict_deriv_at end -/-- If a complex function is differentiable at a real point, then the induced real function is also -differentiable at this point, with a derivative equal to the real part of the complex derivative. -/ +/-- If a complex function `e` is differentiable at a real point, then the function `ℝ → ℝ` given by +the real part of `e` is also differentiable at this point, with a derivative equal to the real part +of the complex derivative. -/ theorem has_deriv_at.real_of_complex (h : has_deriv_at e e' z) : has_deriv_at (λx:ℝ, (e x).re) e'.re z := begin @@ -115,6 +116,19 @@ lemma has_deriv_within_at.complex_to_real_fderiv {f : ℂ → ℂ} {s : set ℂ} by simpa only [complex.restrict_scalars_one_smul_right] using h.has_fderiv_within_at.restrict_scalars ℝ +/-- If a complex function `e` is differentiable at a real point, then its restriction to `ℝ` is +differentiable there as a function `ℝ → ℂ`, with the same derivative. -/ +lemma has_deriv_at.comp_of_real (hf : has_deriv_at e e' ↑z) : has_deriv_at (λ (y:ℝ), e ↑y) e' z := +by simpa only [of_real_clm_apply, of_real_one, mul_one] + using hf.comp z of_real_clm.has_deriv_at + +/-- If a function `f : ℝ → ℝ` is differentiable at a (real) point `x`, then it is also +differentiable as a function `ℝ → ℂ`. -/ +lemma has_deriv_at.of_real_comp {f : ℝ → ℝ} {u : ℝ} (hf : has_deriv_at f u z) : +has_deriv_at (λ (y:ℝ), ↑(f y) : ℝ → ℂ) u z := +by simpa only [of_real_clm_apply, of_real_one, real_smul, mul_one] + using of_real_clm.has_deriv_at.scomp z hf + end real_deriv_of_complex section conformality diff --git a/src/analysis/special_functions/gamma.lean b/src/analysis/special_functions/gamma.lean index a090be162a3f6..e0d175e767c5b 100644 --- a/src/analysis/special_functions/gamma.lean +++ b/src/analysis/special_functions/gamma.lean @@ -118,7 +118,7 @@ def Gamma_integral (s : ℂ) : ℂ := ∫ x in Ioi (0:ℝ), ↑(-x).exp * ↑x ^ lemma Gamma_integral_of_real (s : ℝ) : Gamma_integral ↑s = ↑(s.Gamma_integral) := begin - rw [real.Gamma_integral, ←integral_of_real], + rw [real.Gamma_integral, ←_root_.integral_of_real], refine set_integral_congr measurable_set_Ioi _, intros x hx, dsimp only, rw [of_real_mul, of_real_cpow (mem_Ioi.mp hx).le], @@ -196,14 +196,12 @@ begin { intros x hx, have d1 : has_deriv_at (λ (y: ℝ), (-y).exp) (-(-x).exp) x, { simpa using (has_deriv_at_neg x).exp }, - have d1b : has_deriv_at (λ y, ↑(-y).exp : ℝ → ℂ) (↑-(-x).exp) x, - { convert has_deriv_at.scomp x of_real_clm.has_deriv_at d1, simp, }, - have d2: has_deriv_at (λ (y : ℝ), ↑y ^ s) (s * x ^ (s - 1)) x, - { have t := @has_deriv_at.cpow_const _ _ _ s (has_deriv_at_id ↑x), - simp only [id.def, of_real_re, of_real_im, - ne.def, eq_self_iff_true, not_true, or_false, mul_one] at t, - simpa using has_deriv_at.comp x (t hx.left) of_real_clm.has_deriv_at, }, - simpa only [of_real_neg, neg_mul] using d1b.mul d2 }, + have d2 : has_deriv_at (λ (y : ℝ), ↑y ^ s) (s * x ^ (s - 1)) x, + { have t := @has_deriv_at.cpow_const _ _ _ s (has_deriv_at_id ↑x) _, + simpa only [mul_one] using t.comp_of_real, + simpa only [id.def, of_real_re, of_real_im, + ne.def, eq_self_iff_true, not_true, or_false, mul_one] using hx.1, }, + simpa only [of_real_neg, neg_mul] using d1.of_real_comp.mul d2 }, have cont := (continuous_of_real.comp continuous_neg.exp).mul (continuous_of_real_cpow_const hs), have der_ible := (Gamma_integrand_deriv_integrable_A hs hX).add @@ -214,13 +212,12 @@ begin rw [interval_integral.integral_add (Gamma_integrand_deriv_integrable_A hs hX) (Gamma_integrand_deriv_integrable_B hs hX), interval_integral.integral_neg, neg_add, neg_neg] at int_eval, - replace int_eval := eq_sub_of_add_eq int_eval, - rw [int_eval, sub_neg_eq_add, neg_sub, add_comm, add_sub], + rw [eq_sub_of_add_eq int_eval, sub_neg_eq_add, neg_sub, add_comm, add_sub], simp only [sub_left_inj, add_left_inj], have : (λ x, (-x).exp * (s * x ^ (s - 1)) : ℝ → ℂ) = (λ x, s * (-x).exp * x ^ (s - 1) : ℝ → ℂ), { ext1, ring,}, rw this, - have t := @integral_const_mul (0:ℝ) X volume _ _ s (λ x:ℝ, (-x).exp * x ^ (s - 1)), + have t := @integral_const_mul 0 X volume _ _ s (λ x:ℝ, (-x).exp * x ^ (s - 1)), dsimp at t, rw [←t, of_real_zero, zero_cpow], { rw [mul_zero, add_zero], congr', ext1, ring }, { contrapose! hs, rw [hs, zero_re] } diff --git a/src/analysis/special_functions/integrals.lean b/src/analysis/special_functions/integrals.lean index 9ec2348c984e2..1d09b92712e00 100644 --- a/src/analysis/special_functions/integrals.lean +++ b/src/analysis/special_functions/integrals.lean @@ -242,12 +242,11 @@ end lemma integral_cpow {r : ℂ} (ha : 0 < a) (hb : 0 < b) (hr : r ≠ -1) : ∫ (x : ℝ) in a..b, (x : ℂ) ^ r = (b ^ (r + 1) - a ^ (r + 1)) / (r + 1) := begin - suffices : ∀ x ∈ set.interval a b, has_deriv_at (λ x : ℝ, (x : ℂ) ^ (r + 1) / (r + 1)) (x ^ r) x, - { rw sub_div, - exact integral_eq_sub_of_has_deriv_at this (interval_integrable_cpow ha hb) }, + rw sub_div, + suffices : ∀ x ∈ set.interval a b, has_deriv_at (λ z : ℂ, z ^ (r + 1) / (r + 1)) (x ^ r) x, + { exact integral_eq_sub_of_has_deriv_at + (λ x hx, (this x hx).comp_of_real) (interval_integrable_cpow ha hb) }, intros x hx, - suffices : has_deriv_at (λ z : ℂ, z ^ (r + 1) / (r + 1)) (x ^ r) x, - { simpa using has_deriv_at.comp x this complex.of_real_clm.has_deriv_at }, have hx' : 0 < (x : ℂ).re ∨ (x : ℂ).im ≠ 0, { left, norm_cast, @@ -345,8 +344,7 @@ begin { intro x, conv { congr, skip, rw ←mul_div_cancel (complex.exp (c * x)) hc, }, convert ((complex.has_deriv_at_exp _).comp x _).div_const c using 1, - simpa only [complex.of_real_clm_apply, complex.of_real_one, one_mul, mul_one, mul_comm] using - complex.of_real_clm.has_deriv_at.mul_const c }, + simpa only [mul_one] using ((has_deriv_at_id (x:ℂ)).const_mul _).comp_of_real, }, rw integral_deriv_eq_sub' _ (funext (λ x, (D x).deriv)) (λ x hx, (D x).differentiable_at), { ring_nf }, { apply continuous.continuous_on, continuity,} diff --git a/src/measure_theory/integral/interval_integral.lean b/src/measure_theory/integral/interval_integral.lean index 2fc5ac0ba29c6..2f18a52d68e12 100644 --- a/src/measure_theory/integral/interval_integral.lean +++ b/src/measure_theory/integral/interval_integral.lean @@ -699,6 +699,10 @@ by simp only [interval_integral, measure.restrict_smul, integral_smul_measure, s end basic +lemma integral_of_real {a b : ℝ} {μ : measure ℝ} {f : ℝ → ℝ} : + ∫ x in a..b, (f x : ℂ) ∂μ = ↑(∫ x in a..b, f x ∂μ) := +by simp only [interval_integral, integral_of_real, complex.of_real_sub] + section continuous_linear_map variables {a b : ℝ} {μ : measure ℝ} {f : ℝ → E} From cd9a9326dc14ad6e438e62267c31c66dd680d94e Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9mi=20Bottinelli?= Date: Mon, 26 Dec 2022 14:12:12 +0000 Subject: [PATCH 0109/1291] feat(analysis/bounded_variation): `evariation_on` of precomposition with monotone/antitone map. (#18001) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Rémi Bottinelli --- src/analysis/bounded_variation.lean | 87 +++++++++++++++++++++++++++-- src/data/set/function.lean | 17 ++++++ 2 files changed, 98 insertions(+), 6 deletions(-) diff --git a/src/analysis/bounded_variation.lean b/src/analysis/bounded_variation.lean index b8531f14bb883..23a659043b5cc 100644 --- a/src/analysis/bounded_variation.lean +++ b/src/analysis/bounded_variation.lean @@ -5,6 +5,7 @@ Authors: Sébastien Gouëzel -/ import measure_theory.measure.lebesgue import analysis.calculus.monotone +import data.set.function /-! # Functions of bounded variation @@ -44,7 +45,7 @@ that the sets one uses are nonempty and bounded above as these are only conditio open_locale big_operators nnreal ennreal open set measure_theory -variables {α : Type*} [linear_order α] +variables {α β : Type*} [linear_order α] [linear_order β] {E F : Type*} [pseudo_emetric_space E] [pseudo_emetric_space F] {V : Type*} [normed_add_comm_group V] [normed_space ℝ V] [finite_dimensional ℝ V] @@ -75,6 +76,16 @@ begin exact ⟨⟨λ i, x, λ i j hij, le_rfl, λ i, hx⟩⟩, end +lemma eq_of_eq_on {f f' : α → E} {s : set α} (h : set.eq_on f f' s) : + evariation_on f s = evariation_on f' s := +begin + dsimp only [evariation_on], + congr' 1 with p : 1, + congr' 1 with i : 1, + congr' 1; + exact h (p.2.2.2 _), +end + lemma sum_le (f : α → E) {s : set α} (n : ℕ) {u : ℕ → α} (hu : monotone u) (us : ∀ i, u i ∈ s) : ∑ i in finset.range n, edist (f (u (i+1))) (f (u i)) ≤ evariation_on f s := @@ -167,16 +178,19 @@ lemma _root_.has_bounded_variation_on.has_locally_bounded_variation_on {f : α (h : has_bounded_variation_on f s) : has_locally_bounded_variation_on f s := λ x y hx hy, h.mono (inter_subset_left _ _) -@[simp] protected lemma subsingleton (f : α → E) {s : set α} (hs : s.subsingleton) : - evariation_on f s = 0 := +lemma constant_on {f : α → E} {s : set α} + (hf : (f '' s).subsingleton) : evariation_on f s = 0 := begin apply le_antisymm _ (zero_le _), apply supr_le _, rintros ⟨n, ⟨u, hu, ut⟩⟩, - have : ∀ i, u i = u 0, from λ i, hs (ut _) (ut _), + have : ∀ i, f (u i) = f (u 0) := λ i, hf ⟨u i, ut i, rfl⟩ ⟨u 0, ut 0, rfl⟩, simp [subtype.coe_mk, le_zero_iff, finset.sum_eq_zero_iff, finset.mem_range, this], end +@[simp] protected lemma subsingleton (f : α → E) {s : set α} (hs : s.subsingleton) : + evariation_on f s = 0 := constant_on (hs.image f) + lemma edist_le (f : α → E) {s : set α} {x y : α} (hx : x ∈ s) (hy : y ∈ s) : edist (f x) (f y) ≤ evariation_on f s := begin @@ -520,6 +534,69 @@ begin rw [← evariation_on.union f A B, ← inter_union_distrib_left, Icc_union_Icc_eq_Icc hab hbc], end +lemma comp_le_of_monotone_on (f : α → E) {s : set α} {t : set β} (φ : β → α) + (hφ : monotone_on φ t) (φst : set.maps_to φ t s) : + evariation_on (f ∘ φ) t ≤ evariation_on f s := +begin + apply supr_le _, + rintro ⟨n, ⟨u, hu, ut⟩⟩, + exact le_supr (λ (p : ℕ × {u : ℕ → α // monotone u ∧ ∀ i, u i ∈ s}), + ∑ i in finset.range p.1, edist (f ((p.2 : ℕ → α) (i+1))) (f ((p.2 : ℕ → α) i))) + ⟨n, ⟨φ ∘ u, λ x y xy, hφ (ut x) (ut y) (hu xy), λ i, φst (ut i)⟩⟩, +end + +lemma comp_le_of_antitone_on (f : α → E) {s : set α} {t : set β} (φ : β → α) + (hφ : antitone_on φ t) (φst : set.maps_to φ t s) : + evariation_on (f ∘ φ) t ≤ evariation_on f s := +begin + apply supr_le _, + rintros ⟨n, ⟨u, hu, ut⟩⟩, + change ∑ i in finset.range n, edist (f ∘ φ $ u (i+1)) (f ∘ φ $ u i) ≤ evariation_on f s, + rw ←finset.sum_range_reflect, + have : ∀ x : ℕ, x ∈ finset.range n → + edist ((f ∘ φ) (u (n - 1 - x + 1))) ((f ∘ φ) (u (n - 1 - x))) = + edist ((f ∘ φ) (u (n - (x + 1)))) ((f ∘ φ) (u (n - x))) := λ x hx, by + { rw [edist_comm, nat.sub_sub, add_comm, nat.sub_succ, nat.add_one, nat.succ_pred_eq_of_pos], + simpa only [tsub_pos_iff_lt, finset.mem_range] using hx, }, + rw finset.sum_congr rfl this, + let ru : ℕ → β := λ i, u (n-i), + have rut : ∀ i : ℕ, ru i ∈ t := λ i, ut (n-i), + have hru : antitone ru := λ i j l, hu (n.sub_le_sub_left l), + exact le_supr (λ (p : ℕ × {u : ℕ → α // monotone u ∧ ∀ i, u i ∈ s}), + ∑ i in finset.range p.1, edist (f ((p.2 : ℕ → α) (i+1))) (f ((p.2 : ℕ → α) i))) + ⟨n, ⟨φ ∘ ru, λ x y xy, hφ (rut y) (rut x) (hru xy), λ i, φst (rut i)⟩⟩, +end + +lemma comp_eq_of_monotone_on (f : α → E) {s : set α} {t : set β} [nonempty β] (φ : β → α) + (hφ : monotone_on φ t) (φst : set.maps_to φ t s) (φsur : set.surj_on φ t s) : + evariation_on (f ∘ φ) t = evariation_on f s := +begin + apply le_antisymm (comp_le_of_monotone_on f φ hφ φst), + let ψ := φ.inv_fun_on t, + have ψφs : set.eq_on (φ ∘ ψ) id s := φsur.right_inv_on_inv_fun_on, + have ψts : set.maps_to ψ s t := φsur.maps_to_inv_fun_on, + have hψ : monotone_on ψ s := + function.monotone_on_of_right_inv_on_of_maps_to hφ ψφs ψts, + change evariation_on (f ∘ id) s ≤ evariation_on (f ∘ φ) t, + rw ←eq_of_eq_on (ψφs.comp_left : set.eq_on (f ∘ (φ ∘ ψ)) (f ∘ id) s), + apply comp_le_of_monotone_on _ ψ hψ ψts, +end + +lemma comp_eq_of_antitone_on (f : α → E) {s : set α} {t : set β} [nonempty β] (φ : β → α) + (hφ : antitone_on φ t) (φst : set.maps_to φ t s) (φsur : set.surj_on φ t s) : + evariation_on (f ∘ φ) t = evariation_on f s := +begin + apply le_antisymm (comp_le_of_antitone_on f φ hφ φst), + let ψ := φ.inv_fun_on t, + have ψφs : set.eq_on (φ ∘ ψ) id s := φsur.right_inv_on_inv_fun_on, + have ψts : set.maps_to ψ s t := φsur.maps_to_inv_fun_on, + have hψ : antitone_on ψ s := + function.antitone_on_of_right_inv_on_of_maps_to hφ ψφs ψts, + change evariation_on (f ∘ id) s ≤ evariation_on (f ∘ φ) t, + rw ←eq_of_eq_on (ψφs.comp_left : set.eq_on (f ∘ (φ ∘ ψ)) (f ∘ id) s), + apply comp_le_of_antitone_on _ ψ hψ ψts, +end + end evariation_on /-! ## Monotone functions and bounded variation -/ @@ -756,5 +833,3 @@ lemma lipschitz_with.ae_differentiable_at {C : ℝ≥0} {f : ℝ → V} (h : lipschitz_with C f) : ∀ᵐ x, differentiable_at ℝ f x := (h.has_locally_bounded_variation_on univ).ae_differentiable_at - - diff --git a/src/data/set/function.lean b/src/data/set/function.lean index 07be41cf6a539..642d3f19b3a7b 100644 --- a/src/data/set/function.lean +++ b/src/data/set/function.lean @@ -1286,4 +1286,21 @@ update_comp_eq_of_not_mem_range' g a h lemma insert_inj_on (s : set α) : sᶜ.inj_on (λ a, insert a s) := λ a ha b _, (insert_inj ha).1 +lemma monotone_on_of_right_inv_on_of_maps_to + [partial_order α] [linear_order β] {φ : β → α} {ψ : α → β} {t : set β} {s : set α} + (hφ : monotone_on φ t) (φψs : set.right_inv_on ψ φ s) (ψts : set.maps_to ψ s t) : + monotone_on ψ s := +begin + rintro x xs y ys l, + rcases le_total (ψ x) (ψ y) with (ψxy|ψyx), + { exact ψxy, }, + { cases le_antisymm l (φψs.eq ys ▸ φψs.eq xs ▸ hφ (ψts ys) (ψts xs) ψyx), refl, }, +end + +lemma antitone_on_of_right_inv_on_of_maps_to + [partial_order α] [linear_order β] {φ : β → α} {ψ : α → β} {t : set β} {s : set α} + (hφ : antitone_on φ t) (φψs : set.right_inv_on ψ φ s) (ψts : set.maps_to ψ s t) : + antitone_on ψ s := +(monotone_on_of_right_inv_on_of_maps_to hφ.dual_left φψs ψts).dual_right + end function From 839ab10c4d112f627e14d9e6d8a4faca5b3f5c86 Mon Sep 17 00:00:00 2001 From: Vincent Beffara Date: Mon, 26 Dec 2022 17:52:20 +0000 Subject: [PATCH 0110/1291] feat(analysis/complex/locally_uniform_limit): compact limits of holomorphic fns are holomorphic (#17074) --- docs/undergrad.yaml | 4 +- src/analysis/calculus/diff_cont_on_cl.lean | 5 + .../calculus/uniform_limits_deriv.lean | 22 +-- src/analysis/complex/cauchy_integral.lean | 14 ++ .../complex/locally_uniform_limit.lean | 182 ++++++++++++++++++ .../complex/removable_singularity.lean | 33 +++- 6 files changed, 246 insertions(+), 14 deletions(-) create mode 100644 src/analysis/complex/locally_uniform_limit.lean diff --git a/docs/undergrad.yaml b/docs/undergrad.yaml index d7c5dee0a7c3b..b775dd5c45787 100644 --- a/docs/undergrad.yaml +++ b/docs/undergrad.yaml @@ -379,7 +379,7 @@ Single Variable Complex Analysis: representations of the $\log$ function on $\C$: '' theorem of holomorphic functions under integral domains: '' winding number of a closed curve in $\C$ with respect to a point: '' - Cauchy formulas: '' + Cauchy formulas: 'complex.two_pi_I_inv_smul_circle_integral_sub_inv_smul_of_differentiable_on_off_countable' analyticity of a holomorphic function: 'differentiable_on.analytic_at' principle of isolated zeros: 'analytic_at.eventually_eq_zero_or_eventually_ne_zero' principle of analytic continuation: 'analytic_on.eq_on_of_preconnected_of_frequently_eq' @@ -389,7 +389,7 @@ Single Variable Complex Analysis: meromorphic functions: '' residue theorem: '' sequences and series of holomorphic functions: '' - holomorphic stability under uniform convergence: '' + holomorphic stability under uniform convergence: 'tendsto_locally_uniformly_on.differentiable_on' # 8. Topology: diff --git a/src/analysis/calculus/diff_cont_on_cl.lean b/src/analysis/calculus/diff_cont_on_cl.lean index bcb209415ff0c..9dac40f6a3871 100644 --- a/src/analysis/calculus/diff_cont_on_cl.lean +++ b/src/analysis/calculus/diff_cont_on_cl.lean @@ -129,3 +129,8 @@ lemma differentiable.comp_diff_cont_on_cl {g : G → E} {t : set G} (hf : differentiable 𝕜 f) (hg : diff_cont_on_cl 𝕜 g t) : diff_cont_on_cl 𝕜 (f ∘ g) t := hf.diff_cont_on_cl.comp hg (maps_to_image _ _) + +lemma differentiable_on.diff_cont_on_cl_ball {U : set E} {c : E} {R : ℝ} + (hf : differentiable_on 𝕜 f U) (hc : closed_ball c R ⊆ U) : + diff_cont_on_cl 𝕜 f (ball c R) := +diff_cont_on_cl.mk_ball (hf.mono (ball_subset_closed_ball.trans hc)) (hf.continuous_on.mono hc) diff --git a/src/analysis/calculus/uniform_limits_deriv.lean b/src/analysis/calculus/uniform_limits_deriv.lean index 80141f59049cc..0e1a612cfc533 100644 --- a/src/analysis/calculus/uniform_limits_deriv.lean +++ b/src/analysis/calculus/uniform_limits_deriv.lean @@ -562,16 +562,15 @@ end lemma has_deriv_at_of_tendsto_locally_uniformly_on [ne_bot l] {s : set 𝕜} (hs : is_open s) (hf' : tendsto_locally_uniformly_on f' g' l s) - (hf : ∀ n, ∀ x ∈ s, has_deriv_at (f n) (f' n x) x) + (hf : ∀ᶠ n in l, ∀ x ∈ s, has_deriv_at (f n) (f' n x) x) (hfg : ∀ x ∈ s, tendsto (λ n, f n x) l (𝓝 (g x))) (hx : x ∈ s) : has_deriv_at g (g' x) x := begin have h1 : s ∈ 𝓝 x := hs.mem_nhds hx, - have h3 : set.univ ×ˢ s ∈ l ×ᶠ 𝓝 x := by simp only [h1, prod_mem_prod_iff, univ_mem, and_self], - have h4 : ∀ᶠ (n : ι × 𝕜) in l ×ᶠ 𝓝 x, has_deriv_at (f n.1) (f' n.1 n.2) n.2, - from eventually_of_mem h3 (λ ⟨n, z⟩ ⟨hn, hz⟩, hf n z hz), - refine has_deriv_at_of_tendsto_uniformly_on_filter _ h4 (eventually_of_mem h1 hfg), + have h2 : ∀ᶠ (n : ι × 𝕜) in l ×ᶠ 𝓝 x, has_deriv_at (f n.1) (f' n.1 n.2) n.2, + from eventually_prod_iff.2 ⟨_, hf, λ x, x ∈ s, h1, λ n, id⟩, + refine has_deriv_at_of_tendsto_uniformly_on_filter _ h2 (eventually_of_mem h1 hfg), simpa [is_open.nhds_within_eq hs hx] using tendsto_locally_uniformly_on_iff_filter.mp hf' x hx, end @@ -580,31 +579,32 @@ terms of `differentiable_on` rather than `has_deriv_at`. This makes a few proofs analysis where holomorphicity is assumed but the derivative is not known a priori. -/ lemma has_deriv_at_of_tendsto_locally_uniformly_on' [ne_bot l] {s : set 𝕜} (hs : is_open s) (hf' : tendsto_locally_uniformly_on (deriv ∘ f) g' l s) - (hf : ∀ n, differentiable_on 𝕜 (f n) s) + (hf : ∀ᶠ n in l, differentiable_on 𝕜 (f n) s) (hfg : ∀ x ∈ s, tendsto (λ n, f n x) l (𝓝 (g x))) (hx : x ∈ s) : has_deriv_at g (g' x) x := begin - refine has_deriv_at_of_tendsto_locally_uniformly_on hs hf' (λ n z hz, _) hfg hx, - exact ((hf n z hz).differentiable_at (hs.mem_nhds hz)).has_deriv_at + refine has_deriv_at_of_tendsto_locally_uniformly_on hs hf' _ hfg hx, + filter_upwards [hf] with n h z hz using ((h z hz).differentiable_at (hs.mem_nhds hz)).has_deriv_at end lemma has_deriv_at_of_tendsto_uniformly_on [ne_bot l] {s : set 𝕜} (hs : is_open s) (hf' : tendsto_uniformly_on f' g' l s) - (hf : ∀ (n : ι), ∀ (x : 𝕜), x ∈ s → has_deriv_at (f n) (f' n x) x) + (hf : ∀ᶠ n in l, ∀ (x : 𝕜), x ∈ s → has_deriv_at (f n) (f' n x) x) (hfg : ∀ (x : 𝕜), x ∈ s → tendsto (λ n, f n x) l (𝓝 (g x))) : ∀ (x : 𝕜), x ∈ s → has_deriv_at g (g' x) x := λ x, has_deriv_at_of_tendsto_locally_uniformly_on hs hf'.tendsto_locally_uniformly_on hf hfg lemma has_deriv_at_of_tendsto_uniformly [ne_bot l] (hf' : tendsto_uniformly f' g' l) - (hf : ∀ (n : ι), ∀ (x : 𝕜), has_deriv_at (f n) (f' n x) x) + (hf : ∀ᶠ n in l, ∀ (x : 𝕜), has_deriv_at (f n) (f' n x) x) (hfg : ∀ (x : 𝕜), tendsto (λ n, f n x) l (𝓝 (g x))) : ∀ (x : 𝕜), has_deriv_at g (g' x) x := begin intros x, - have hf : ∀ (n : ι), ∀ (x : 𝕜), x ∈ set.univ → has_deriv_at (f n) (f' n x) x, { simp [hf], }, + have hf : ∀ᶠ n in l, ∀ (x : 𝕜), x ∈ set.univ → has_deriv_at (f n) (f' n x) x, + by filter_upwards [hf] with n h x hx using h x, have hfg : ∀ (x : 𝕜), x ∈ set.univ → tendsto (λ n, f n x) l (𝓝 (g x)), { simp [hfg], }, have hf' : tendsto_uniformly_on f' g' l set.univ, { rwa tendsto_uniformly_on_univ, }, exact has_deriv_at_of_tendsto_uniformly_on is_open_univ hf' hf hfg x (set.mem_univ x), diff --git a/src/analysis/complex/cauchy_integral.lean b/src/analysis/complex/cauchy_integral.lean index 18fe7a8917912..d1605f7aad8f7 100644 --- a/src/analysis/complex/cauchy_integral.lean +++ b/src/analysis/complex/cauchy_integral.lean @@ -503,6 +503,20 @@ lemma _root_.diff_cont_on_cl.circle_integral_sub_inv_smul {R : ℝ} {c w : ℂ} circle_integral_sub_inv_smul_of_differentiable_on_off_countable countable_empty hw h.continuous_on_ball $ λ x hx, h.differentiable_at is_open_ball hx.1 +/-- **Cauchy integral formula**: if `f : ℂ → E` is complex differentiable on an open disc and is +continuous on its closure, then for any `w` in this open ball we have +$\frac{1}{2πi}\oint_{|z-c|=R}(z-w)^{-1}f(z)\,dz=f(w)$. -/ +lemma _root_.diff_cont_on_cl.two_pi_I_inv_smul_circle_integral_sub_inv_smul {R : ℝ} {c w : ℂ} + {f : ℂ → E} (hf : diff_cont_on_cl ℂ f (ball c R)) (hw : w ∈ ball c R) : + (2 * π * I : ℂ)⁻¹ • ∮ z in C(c, R), (z - w)⁻¹ • f z = f w := +begin + have hR : 0 < R := not_le.mp (ball_eq_empty.not.mp (nonempty_of_mem hw).ne_empty), + refine two_pi_I_inv_smul_circle_integral_sub_inv_smul_of_differentiable_on_off_countable + countable_empty hw _ _, + { simpa only [closure_ball c hR.ne.symm] using hf.continuous_on }, + { simpa only [diff_empty] using λ z hz, hf.differentiable_at is_open_ball hz } +end + /-- **Cauchy integral formula**: if `f : ℂ → E` is complex differentiable on a closed disc of radius `R`, then for any `w` in its interior we have $\oint_{|z-c|=R}(z-w)^{-1}f(z)\,dz=2πif(w)$. -/ lemma _root_.differentiable_on.circle_integral_sub_inv_smul {R : ℝ} {c w : ℂ} {f : ℂ → E} diff --git a/src/analysis/complex/locally_uniform_limit.lean b/src/analysis/complex/locally_uniform_limit.lean new file mode 100644 index 0000000000000..549d05693862d --- /dev/null +++ b/src/analysis/complex/locally_uniform_limit.lean @@ -0,0 +1,182 @@ +/- +Copyright (c) 2022 Vincent Beffara. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Vincent Beffara +-/ +import analysis.complex.removable_singularity +import analysis.calculus.uniform_limits_deriv + +/-! +# Locally uniform limits of holomorphic functions + +This file gathers some results about locally uniform limits of holomorphic functions on an open +subset of the complex plane. + +## Main results + +* `tendsto_locally_uniformly_on.differentiable_on`: A locally uniform limit of holomorphic functions + is holomorphic. +* `tendsto_locally_uniformly_on.deriv`: Locally uniform convergence implies locally uniform + convergence of the derivatives to the derivative of the limit. +-/ + +open set metric measure_theory filter complex interval_integral +open_locale real topological_space + +variables {E ι : Type*} [normed_add_comm_group E] [normed_space ℂ E] [complete_space E] + {U K : set ℂ} {z : ℂ} {M r δ : ℝ} {φ : filter ι} {F : ι → ℂ → E} {f g : ℂ → E} + +namespace complex + +section cderiv + +/-- A circle integral which coincides with `deriv f z` whenever one can apply the Cauchy formula for +the derivative. It is useful in the proof that locally uniform limits of holomorphic functions are +holomorphic, because it depends continuously on `f` for the uniform topology. -/ +noncomputable def cderiv (r : ℝ) (f : ℂ → E) (z : ℂ) : E := +(2 * π * I : ℂ)⁻¹ • ∮ w in C(z, r), ((w - z) ^ 2)⁻¹ • f w + +lemma cderiv_eq_deriv (hU : is_open U) (hf : differentiable_on ℂ f U) (hr : 0 < r) + (hzr : closed_ball z r ⊆ U) : + cderiv r f z = deriv f z := +two_pi_I_inv_smul_circle_integral_sub_sq_inv_smul_of_differentiable hU hzr hf (mem_ball_self hr) + +lemma norm_cderiv_le (hr : 0 < r) (hf : ∀ w ∈ sphere z r, ‖f w‖ ≤ M) : + ‖cderiv r f z‖ ≤ M / r := +begin + have hM : 0 ≤ M, + { obtain ⟨w, hw⟩ : (sphere z r).nonempty := normed_space.sphere_nonempty.mpr hr.le, + exact (norm_nonneg _).trans (hf w hw) }, + have h1 : ∀ w ∈ sphere z r, ‖((w - z) ^ 2)⁻¹ • f w‖ ≤ M / r ^ 2, + { intros w hw, + simp only [mem_sphere_iff_norm, norm_eq_abs] at hw, + simp only [norm_smul, inv_mul_eq_div, hw, norm_eq_abs, map_inv₀, complex.abs_pow], + exact div_le_div hM (hf w hw) (sq_pos_of_pos hr) le_rfl }, + have h2 := circle_integral.norm_integral_le_of_norm_le_const hr.le h1, + simp only [cderiv, norm_smul], + refine (mul_le_mul le_rfl h2 (norm_nonneg _) (norm_nonneg _)).trans (le_of_eq _), + field_simp [_root_.abs_of_nonneg real.pi_pos.le, real.pi_pos.ne.symm, hr.ne.symm], + ring +end + +lemma cderiv_sub (hr : 0 < r) (hf : continuous_on f (sphere z r)) + (hg : continuous_on g (sphere z r)) : + cderiv r (f - g) z = cderiv r f z - cderiv r g z := +begin + have h1 : continuous_on (λ (w : ℂ), ((w - z) ^ 2)⁻¹) (sphere z r), + { refine ((continuous_id'.sub continuous_const).pow 2).continuous_on.inv₀ (λ w hw h, hr.ne _), + rwa [mem_sphere_iff_norm, sq_eq_zero_iff.mp h, norm_zero] at hw }, + simp_rw [cderiv, ← smul_sub], + congr' 1, + simpa only [pi.sub_apply, smul_sub] using circle_integral.integral_sub + ((h1.smul hf).circle_integrable hr.le) ((h1.smul hg).circle_integrable hr.le) +end + +lemma norm_cderiv_lt (hr : 0 < r) (hfM : ∀ w ∈ sphere z r, ‖f w‖ < M) + (hf : continuous_on f (sphere z r)) : + ‖cderiv r f z‖ < M / r := +begin + obtain ⟨L, hL1, hL2⟩ : ∃ L < M, ∀ w ∈ sphere z r, ‖f w‖ ≤ L, + { have e1 : (sphere z r).nonempty := normed_space.sphere_nonempty.mpr hr.le, + have e2 : continuous_on (λ w, ‖f w‖) (sphere z r), + from continuous_norm.comp_continuous_on hf, + obtain ⟨x, hx, hx'⟩ := (is_compact_sphere z r).exists_forall_ge e1 e2, + exact ⟨‖f x‖, hfM x hx, hx'⟩ }, + exact (norm_cderiv_le hr hL2).trans_lt ((div_lt_div_right hr).mpr hL1) +end + +lemma norm_cderiv_sub_lt (hr : 0 < r) (hfg : ∀ w ∈ sphere z r, ‖f w - g w‖ < M) + (hf : continuous_on f (sphere z r)) (hg : continuous_on g (sphere z r)) : + ‖cderiv r f z - cderiv r g z‖ < M / r := +cderiv_sub hr hf hg ▸ norm_cderiv_lt hr hfg (hf.sub hg) + +lemma tendsto_uniformly_on.cderiv (hF : tendsto_uniformly_on F f φ (cthickening δ K)) (hδ : 0 < δ) + (hFn : ∀ᶠ n in φ, continuous_on (F n) (cthickening δ K)) : + tendsto_uniformly_on (cderiv δ ∘ F) (cderiv δ f) φ K := +begin + by_cases φ = ⊥, + { simp only [h, tendsto_uniformly_on, eventually_bot, implies_true_iff]}, + haveI : φ.ne_bot := ne_bot_iff.2 h, + have e1 : continuous_on f (cthickening δ K) := tendsto_uniformly_on.continuous_on hF hFn, + rw [tendsto_uniformly_on_iff] at hF ⊢, + rintro ε hε, + filter_upwards [hF (ε * δ) (mul_pos hε hδ), hFn] with n h h' z hz, + simp_rw [dist_eq_norm] at h ⊢, + have e2 : ∀ w ∈ sphere z δ, ‖f w - F n w‖ < ε * δ, + from λ w hw1, h w (closed_ball_subset_cthickening hz δ (sphere_subset_closed_ball hw1)), + have e3 := sphere_subset_closed_ball.trans (closed_ball_subset_cthickening hz δ), + have hf : continuous_on f (sphere z δ), + from e1.mono (sphere_subset_closed_ball.trans (closed_ball_subset_cthickening hz δ)), + simpa only [mul_div_cancel _ hδ.ne.symm] using norm_cderiv_sub_lt hδ e2 hf (h'.mono e3) +end + +end cderiv + +section weierstrass + +lemma tendsto_uniformly_on_deriv_of_cthickening_subset (hf : tendsto_locally_uniformly_on F f φ U) + (hF : ∀ᶠ n in φ, differentiable_on ℂ (F n) U) {δ : ℝ} (hδ: 0 < δ) (hK : is_compact K) + (hU : is_open U) (hKU : cthickening δ K ⊆ U) : + tendsto_uniformly_on (deriv ∘ F) (cderiv δ f) φ K := +begin + have h1 : ∀ᶠ n in φ, continuous_on (F n) (cthickening δ K), + by filter_upwards [hF] with n h using h.continuous_on.mono hKU, + have h2 : is_compact (cthickening δ K), + from is_compact_of_is_closed_bounded is_closed_cthickening hK.bounded.cthickening, + have h3 : tendsto_uniformly_on F f φ (cthickening δ K), + from (tendsto_locally_uniformly_on_iff_forall_is_compact hU).mp hf (cthickening δ K) hKU h2, + apply (h3.cderiv hδ h1).congr, + filter_upwards [hF] with n h z hz, + exact cderiv_eq_deriv hU h hδ ((closed_ball_subset_cthickening hz δ).trans hKU) +end + +lemma exists_cthickening_tendsto_uniformly_on (hf : tendsto_locally_uniformly_on F f φ U) + (hF : ∀ᶠ n in φ, differentiable_on ℂ (F n) U) (hK : is_compact K) (hU : is_open U) (hKU : K ⊆ U) : + ∃ δ > 0, cthickening δ K ⊆ U ∧ tendsto_uniformly_on (deriv ∘ F) (cderiv δ f) φ K := +begin + obtain ⟨δ, hδ, hKδ⟩ := hK.exists_cthickening_subset_open hU hKU, + exact ⟨δ, hδ, hKδ, tendsto_uniformly_on_deriv_of_cthickening_subset hf hF hδ hK hU hKδ⟩ +end + +/-- A locally uniform limit of holomorphic functions on an open domain of the complex plane is +holomorphic (the derivatives converge locally uniformly to that of the limit, which is proved +as `tendsto_locally_uniformly_on.deriv`). -/ +theorem _root_.tendsto_locally_uniformly_on.differentiable_on [φ.ne_bot] + (hf : tendsto_locally_uniformly_on F f φ U) (hF : ∀ᶠ n in φ, differentiable_on ℂ (F n) U) + (hU : is_open U) : + differentiable_on ℂ f U := +begin + rintro x hx, + obtain ⟨K, ⟨hKx, hK⟩, hKU⟩ := (compact_basis_nhds x).mem_iff.mp (hU.mem_nhds hx), + obtain ⟨δ, hδ, -, h1⟩ := exists_cthickening_tendsto_uniformly_on hf hF hK hU hKU, + have h2 : interior K ⊆ U := interior_subset.trans hKU, + have h3 : ∀ᶠ n in φ, differentiable_on ℂ (F n) (interior K), + filter_upwards [hF] with n h using h.mono h2, + have h4 : tendsto_locally_uniformly_on F f φ (interior K) := hf.mono h2, + have h5 : tendsto_locally_uniformly_on (deriv ∘ F) (cderiv δ f) φ (interior K), + from h1.tendsto_locally_uniformly_on.mono interior_subset, + have h6 : ∀ x ∈ interior K, has_deriv_at f (cderiv δ f x) x, + from λ x h, has_deriv_at_of_tendsto_locally_uniformly_on' + is_open_interior h5 h3 (λ _, h4.tendsto_at) h, + have h7 : differentiable_on ℂ f (interior K), + from λ x hx, (h6 x hx).differentiable_at.differentiable_within_at, + exact (h7.differentiable_at (interior_mem_nhds.mpr hKx)).differentiable_within_at +end + +lemma _root_.tendsto_locally_uniformly_on.deriv (hf : tendsto_locally_uniformly_on F f φ U) + (hF : ∀ᶠ n in φ, differentiable_on ℂ (F n) U) (hU : is_open U) : + tendsto_locally_uniformly_on (deriv ∘ F) (deriv f) φ U := +begin + rw [tendsto_locally_uniformly_on_iff_forall_is_compact hU], + by_cases φ = ⊥, + { simp only [h, tendsto_uniformly_on, eventually_bot, implies_true_iff] }, + haveI : φ.ne_bot := ne_bot_iff.2 h, + rintro K hKU hK, + obtain ⟨δ, hδ, hK4, h⟩ := exists_cthickening_tendsto_uniformly_on hf hF hK hU hKU, + refine h.congr_right (λ z hz, cderiv_eq_deriv hU (hf.differentiable_on hF hU) hδ _), + exact (closed_ball_subset_cthickening hz δ).trans hK4, +end + +end weierstrass + +end complex diff --git a/src/analysis/complex/removable_singularity.lean b/src/analysis/complex/removable_singularity.lean index ce0694306c384..8ce0db71f46c4 100644 --- a/src/analysis/complex/removable_singularity.lean +++ b/src/analysis/complex/removable_singularity.lean @@ -17,7 +17,7 @@ function `function.update f c (lim (𝓝[≠] c) f)` is complex differentiable i -/ open topological_space metric set filter asymptotics function -open_locale topological_space filter nnreal +open_locale topological_space filter nnreal real universe u variables {E : Type u} [normed_add_comm_group E] [normed_space ℂ E] [complete_space E] @@ -126,4 +126,35 @@ lemma tendsto_lim_of_differentiable_on_punctured_nhds_of_bounded_under {f : ℂ tendsto f (𝓝[≠] c) (𝓝 $ lim (𝓝[≠] c) f) := tendsto_lim_of_differentiable_on_punctured_nhds_of_is_o hd hb.is_o_sub_self_inv +/-- The Cauchy formula for the derivative of a holomorphic function. -/ +lemma two_pi_I_inv_smul_circle_integral_sub_sq_inv_smul_of_differentiable + {U : set ℂ} (hU : is_open U) {c w₀ : ℂ} {R : ℝ} {f : ℂ → E} + (hc : closed_ball c R ⊆ U) (hf : differentiable_on ℂ f U) (hw₀ : w₀ ∈ ball c R) : + (2 * π * I : ℂ)⁻¹ • ∮ z in C(c, R), ((z - w₀) ^ 2)⁻¹ • f z = deriv f w₀ := +begin + -- We apply the removable singularity theorem and the Cauchy formula to `dslope f w₀` + have hR : 0 < R := not_le.mp (ball_eq_empty.not.mp (nonempty_of_mem hw₀).ne_empty), + have hf' : differentiable_on ℂ (dslope f w₀) U, + from (differentiable_on_dslope (hU.mem_nhds ((ball_subset_closed_ball.trans hc) hw₀))).mpr hf, + have h0 := (hf'.diff_cont_on_cl_ball hc).two_pi_I_inv_smul_circle_integral_sub_inv_smul hw₀, + rw [← dslope_same, ← h0], + congr' 1, + transitivity ∮ z in C(c, R), ((z - w₀) ^ 2)⁻¹ • (f z - f w₀), + { have h1 : continuous_on (λ (z : ℂ), ((z - w₀) ^ 2)⁻¹) (sphere c R), + { refine ((continuous_id'.sub continuous_const).pow 2).continuous_on.inv₀ (λ w hw h, _), + exact sphere_disjoint_ball.ne_of_mem hw hw₀ (sub_eq_zero.mp (sq_eq_zero_iff.mp h)) }, + have h2 : circle_integrable (λ (z : ℂ), ((z - w₀) ^ 2)⁻¹ • f z) c R, + { refine continuous_on.circle_integrable (pos_of_mem_ball hw₀).le _, + exact h1.smul (hf.continuous_on.mono (sphere_subset_closed_ball.trans hc)) }, + have h3 : circle_integrable (λ (z : ℂ), ((z - w₀) ^ 2)⁻¹ • f w₀) c R, + from continuous_on.circle_integrable (pos_of_mem_ball hw₀).le (h1.smul continuous_on_const), + have h4 : ∮ (z : ℂ) in C(c, R), ((z - w₀) ^ 2)⁻¹ = 0, + by simpa using circle_integral.integral_sub_zpow_of_ne (dec_trivial : (-2 : ℤ) ≠ -1) c w₀ R, + simp only [smul_sub, circle_integral.integral_sub h2 h3, h4, + circle_integral.integral_smul_const, zero_smul, sub_zero] }, + { refine circle_integral.integral_congr (pos_of_mem_ball hw₀).le (λ z hz, _), + simp only [dslope_of_ne, metric.sphere_disjoint_ball.ne_of_mem hz hw₀, slope, ← smul_assoc, sq, + mul_inv, ne.def, not_false_iff, vsub_eq_sub, algebra.id.smul_eq_mul] } +end + end complex From a7b9272580d5bbcf307c4cba67ca648065c36218 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 26 Dec 2022 19:01:00 +0000 Subject: [PATCH 0111/1291] feat(order/locally_finite): `finset.interval` (#17939) A `finset`-valued version of `set.interval`. --- src/data/finset/locally_finite.lean | 101 +++++++++++++++++++++++++++- src/order/locally_finite.lean | 39 +++++++++++ 2 files changed, 139 insertions(+), 1 deletion(-) diff --git a/src/data/finset/locally_finite.lean b/src/data/finset/locally_finite.lean index 5e12342c85f88..c41563436e776 100644 --- a/src/data/finset/locally_finite.lean +++ b/src/data/finset/locally_finite.lean @@ -23,7 +23,8 @@ https://github.com/leanprover-community/mathlib/pull/14448#discussion_r906109235 for some ideas. -/ -open_locale big_operators +open function order_dual +open_locale big_operators finset_interval variables {ι α : Type*} @@ -498,6 +499,104 @@ by { ext, simp [eq_comm] } end linear_order +section lattice +variables [lattice α] [locally_finite_order α] {a a₁ a₂ b b₁ b₂ c x : α} + +lemma interval_to_dual (a b : α) : [to_dual a, to_dual b] = [a, b].map to_dual.to_embedding := +Icc_to_dual _ _ + +@[simp] lemma interval_of_le (h : a ≤ b) : [a, b] = Icc a b := +by rw [interval, inf_eq_left.2 h, sup_eq_right.2 h] + +@[simp] lemma interval_of_ge (h : b ≤ a) : [a, b] = Icc b a := +by rw [interval, inf_eq_right.2 h, sup_eq_left.2 h] + +lemma interval_swap (a b : α) : [a, b] = [b, a] := by rw [interval, interval, inf_comm, sup_comm] + +@[simp] lemma interval_self : [a, a] = {a} := by simp [interval] + +@[simp] lemma nonempty_interval : finset.nonempty [a, b] := nonempty_Icc.2 inf_le_sup + +lemma Icc_subset_interval : Icc a b ⊆ [a, b] := Icc_subset_Icc inf_le_left le_sup_right +lemma Icc_subset_interval' : Icc b a ⊆ [a, b] := Icc_subset_Icc inf_le_right le_sup_left + +@[simp] lemma left_mem_interval : a ∈ [a, b] := mem_Icc.2 ⟨inf_le_left, le_sup_left⟩ +@[simp] lemma right_mem_interval : b ∈ [a, b] := mem_Icc.2 ⟨inf_le_right, le_sup_right⟩ + +lemma mem_interval_of_le (ha : a ≤ x) (hb : x ≤ b) : x ∈ [a, b] := +Icc_subset_interval $ mem_Icc.2 ⟨ha, hb⟩ + +lemma mem_interval_of_ge (hb : b ≤ x) (ha : x ≤ a) : x ∈ [a, b] := +Icc_subset_interval' $ mem_Icc.2 ⟨hb, ha⟩ + +lemma interval_subset_interval (h₁ : a₁ ∈ [a₂, b₂]) (h₂ : b₁ ∈ [a₂, b₂]) : [a₁, b₁] ⊆ [a₂, b₂] := +by { rw mem_interval at h₁ h₂, exact Icc_subset_Icc (le_inf h₁.1 h₂.1) (sup_le h₁.2 h₂.2) } + +lemma interval_subset_Icc (ha : a₁ ∈ Icc a₂ b₂) (hb : b₁ ∈ Icc a₂ b₂) : [a₁, b₁] ⊆ Icc a₂ b₂ := +by { rw mem_Icc at ha hb, exact Icc_subset_Icc (le_inf ha.1 hb.1) (sup_le ha.2 hb.2) } + +lemma interval_subset_interval_iff_mem : [a₁, b₁] ⊆ [a₂, b₂] ↔ a₁ ∈ [a₂, b₂] ∧ b₁ ∈ [a₂, b₂] := +⟨λ h, ⟨h left_mem_interval, h right_mem_interval⟩, λ h, interval_subset_interval h.1 h.2⟩ + +lemma interval_subset_interval_iff_le' : + [a₁, b₁] ⊆ [a₂, b₂] ↔ a₂ ⊓ b₂ ≤ a₁ ⊓ b₁ ∧ a₁ ⊔ b₁ ≤ a₂ ⊔ b₂ := +Icc_subset_Icc_iff inf_le_sup + +lemma interval_subset_interval_right (h : x ∈ [a, b]) : [x, b] ⊆ [a, b] := +interval_subset_interval h right_mem_interval + +lemma interval_subset_interval_left (h : x ∈ [a, b]) : [a, x] ⊆ [a, b] := +interval_subset_interval left_mem_interval h + +end lattice + +section distrib_lattice +variables [distrib_lattice α] [locally_finite_order α] {a a₁ a₂ b b₁ b₂ c x : α} + +lemma eq_of_mem_interval_of_mem_interval : a ∈ [b, c] → b ∈ [a, c] → a = b := +by { simp_rw mem_interval, exact set.eq_of_mem_interval_of_mem_interval } + +lemma eq_of_mem_interval_of_mem_interval' : b ∈ [a, c] → c ∈ [a, b] → b = c := +by { simp_rw mem_interval, exact set.eq_of_mem_interval_of_mem_interval' } + +lemma interval_injective_right (a : α) : injective (λ b, [b, a]) := +λ b c h, by { rw ext_iff at h, + exact eq_of_mem_interval_of_mem_interval ((h _).1 left_mem_interval) ((h _).2 left_mem_interval) } + +lemma interval_injective_left (a : α) : injective (interval a) := +by simpa only [interval_swap] using interval_injective_right a + +end distrib_lattice + +section linear_order +variables [linear_order α] [locally_finite_order α] {a a₁ a₂ b b₁ b₂ c x : α} + +lemma Icc_min_max : Icc (min a b) (max a b) = [a, b] := rfl + +lemma interval_of_not_le (h : ¬ a ≤ b) : [a, b] = Icc b a := interval_of_ge $ le_of_not_ge h +lemma interval_of_not_ge (h : ¬ b ≤ a) : [a, b] = Icc a b := interval_of_le $ le_of_not_ge h + +lemma interval_eq_union : [a, b] = Icc a b ∪ Icc b a := +coe_injective $ by { push_cast, exact set.interval_eq_union } + +lemma mem_interval' : a ∈ [b, c] ↔ b ≤ a ∧ a ≤ c ∨ c ≤ a ∧ a ≤ b := by simp [interval_eq_union] + +lemma not_mem_interval_of_lt : c < a → c < b → c ∉ [a, b] := +by { rw mem_interval, exact set.not_mem_interval_of_lt } + +lemma not_mem_interval_of_gt : a < c → b < c → c ∉ [a, b] := +by { rw mem_interval, exact set.not_mem_interval_of_gt } + +lemma interval_subset_interval_iff_le : + [a₁, b₁] ⊆ [a₂, b₂] ↔ min a₂ b₂ ≤ min a₁ b₁ ∧ max a₁ b₁ ≤ max a₂ b₂ := +interval_subset_interval_iff_le' + +/-- A sort of triangle inequality. -/ +lemma interval_subset_interval_union_interval : [a, c] ⊆ [a, b] ∪ [b, c] := +coe_subset.1 $ by { push_cast, exact set.interval_subset_interval_union_interval } + +end linear_order + section ordered_cancel_add_comm_monoid variables [ordered_cancel_add_comm_monoid α] [has_exists_add_of_le α] [locally_finite_order α] diff --git a/src/order/locally_finite.lean b/src/order/locally_finite.lean index 8f4ee3a9b7c29..037ce5d44d580 100644 --- a/src/order/locally_finite.lean +++ b/src/order/locally_finite.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ import data.finset.preimage +import data.set.intervals.unordered_interval /-! # Locally finite orders @@ -29,6 +30,7 @@ In a `locally_finite_order`, * `finset.Ico`: Closed-open interval as a finset. * `finset.Ioc`: Open-closed interval as a finset. * `finset.Ioo`: Open-open interval as a finset. +* `finset.interval`: Unordered closed interval as a finset. * `multiset.Icc`: Closed-closed interval as a multiset. * `multiset.Ico`: Closed-open interval as a multiset. * `multiset.Ioc`: Open-closed interval as a multiset. @@ -249,6 +251,7 @@ protected def _root_.is_empty.to_locally_finite_order_bot [preorder α] [is_empt /-! ### Intervals as finsets -/ namespace finset +section preorder variables [preorder α] section locally_finite_order @@ -359,6 +362,24 @@ lemma Iic_eq_Icc : Iic = Icc (⊥ : α) := rfl lemma Iio_eq_Ico : Iio = Ico (⊥ : α) := rfl end order_bot +end preorder + +section lattice +variables [lattice α] [locally_finite_order α] {a b x : α} + +/-- `finset.interval a b` is the set of elements lying between `a` and `b`, with `a` and `b` +included. Note that we define it more generally in a lattice as `finset.Icc (a ⊓ b) (a ⊔ b)`. In a +product type, `finset.interval` corresponds to the bounding box of the two elements. -/ +def interval (a b : α) : finset α := Icc (a ⊓ b) (a ⊔ b) + +localized "notation (name := finset.interval) `[`a `, ` b `]` := finset.interval a b" + in finset_interval + +@[simp] lemma mem_interval : x ∈ interval a b ↔ a ⊓ b ≤ x ∧ x ≤ a ⊔ b := mem_Icc + +@[simp, norm_cast] lemma coe_interval (a b : α) : ([a, b] : set α) = set.interval a b := coe_Icc _ _ + +end lattice end finset /-! ### Intervals as multisets -/ @@ -694,6 +715,24 @@ end prod end preorder +namespace prod +variables [lattice α] [lattice β] + +lemma interval_eq [locally_finite_order α] [locally_finite_order β] + [decidable_rel ((≤) : α × β → α × β → Prop)] (p q : α × β) : + finset.interval p q = finset.interval p.1 q.1 ×ˢ finset.interval p.2 q.2 := rfl + +@[simp] lemma interval_mk_mk [locally_finite_order α] [locally_finite_order β] + [decidable_rel ((≤) : α × β → α × β → Prop)] (a₁ a₂ : α) (b₁ b₂ : β) : + finset.interval (a₁, b₁) (a₂, b₂) = finset.interval a₁ a₂ ×ˢ finset.interval b₁ b₂ := rfl + +lemma card_interval [locally_finite_order α] [locally_finite_order β] + [decidable_rel ((≤) : α × β → α × β → Prop)] (p q : α × β) : + (finset.interval p q).card = (finset.interval p.1 q.1).card * (finset.interval p.2 q.2).card := +prod.card_Icc _ _ + +end prod + /-! #### `with_top`, `with_bot` From 79ceafba6a55c014a7a7299d8c3ef8ead13a6036 Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Tue, 27 Dec 2022 07:44:52 +0000 Subject: [PATCH 0112/1291] perf(analysis/inner_product_space/basic): speedups (#18018) Speed up all 30 declaration with elaboration time > 10s (on my machine) mostly by squeezing simps, except ``` elaboration of abs_inner_div_norm_mul_norm_eq_one_iff took 13.1s ``` is not sped up because I couldn't find a good way. --- src/analysis/inner_product_space/basic.lean | 116 ++++++++++---------- 1 file changed, 60 insertions(+), 56 deletions(-) diff --git a/src/analysis/inner_product_space/basic.lean b/src/analysis/inner_product_space/basic.lean index 147a76d74b76f..79de0faad1c92 100644 --- a/src/analysis/inner_product_space/basic.lean +++ b/src/analysis/inner_product_space/basic.lean @@ -449,24 +449,26 @@ lemma inner_sum {ι : Type*} (s : finset ι) (f : ι → E) (x : E) : lemma finsupp.sum_inner {ι : Type*} (l : ι →₀ 𝕜) (v : ι → E) (x : E) : ⟪l.sum (λ (i : ι) (a : 𝕜), a • v i), x⟫ = l.sum (λ (i : ι) (a : 𝕜), (conj a) • ⟪v i, x⟫) := -by { convert sum_inner l.support (λ a, l a • v a) x, simp [inner_smul_left, finsupp.sum] } +by { convert sum_inner l.support (λ a, l a • v a) x, + simp only [inner_smul_left, finsupp.sum, smul_eq_mul] } /-- An inner product with a sum on the right, `finsupp` version. -/ lemma finsupp.inner_sum {ι : Type*} (l : ι →₀ 𝕜) (v : ι → E) (x : E) : ⟪x, l.sum (λ (i : ι) (a : 𝕜), a • v i)⟫ = l.sum (λ (i : ι) (a : 𝕜), a • ⟪x, v i⟫) := -by { convert inner_sum l.support (λ a, l a • v a) x, simp [inner_smul_right, finsupp.sum] } +by { convert inner_sum l.support (λ a, l a • v a) x, + simp only [inner_smul_right, finsupp.sum, smul_eq_mul] } lemma dfinsupp.sum_inner {ι : Type*} [dec : decidable_eq ι] {α : ι → Type*} [Π i, add_zero_class (α i)] [Π i (x : α i), decidable (x ≠ 0)] (f : Π i, α i → E) (l : Π₀ i, α i) (x : E) : ⟪l.sum f, x⟫ = l.sum (λ i a, ⟪f i a, x⟫) := -by simp [dfinsupp.sum, sum_inner] {contextual := tt} +by simp only [dfinsupp.sum, sum_inner, smul_eq_mul] {contextual := tt} lemma dfinsupp.inner_sum {ι : Type*} [dec : decidable_eq ι] {α : ι → Type*} [Π i, add_zero_class (α i)] [Π i (x : α i), decidable (x ≠ 0)] (f : Π i, α i → E) (l : Π₀ i, α i) (x : E) : ⟪x, l.sum f⟫ = l.sum (λ i a, ⟪x, f i a⟫) := -by simp [dfinsupp.sum, inner_sum] {contextual := tt} +by simp only [dfinsupp.sum, inner_sum, smul_eq_mul] {contextual := tt} @[simp] lemma inner_zero_left {x : E} : ⟪0, x⟫ = 0 := by rw [← zero_smul 𝕜 (0:E), inner_smul_left, ring_hom.map_zero, zero_mul] @@ -514,12 +516,12 @@ lemma real_inner_self_nonpos {x : F} : ⟪x, x⟫_ℝ ≤ 0 ↔ x = 0 := by { have h := @inner_self_nonpos ℝ F _ _ x, simpa using h } @[simp] lemma inner_self_re_to_K {x : E} : (re ⟪x, x⟫ : 𝕜) = ⟪x, x⟫ := -by rw is_R_or_C.ext_iff; exact ⟨by simp, by simp [inner_self_nonneg_im]⟩ +is_R_or_C.ext_iff.2 ⟨by simp only [of_real_re], by simp only [inner_self_nonneg_im, of_real_im]⟩ lemma inner_self_eq_norm_sq_to_K (x : E) : ⟪x, x⟫ = (‖x‖ ^ 2 : 𝕜) := begin suffices : (is_R_or_C.re ⟪x, x⟫ : 𝕜) = ‖x‖ ^ 2, - { simpa [inner_self_re_to_K] using this }, + { simpa only [inner_self_re_to_K] using this }, exact_mod_cast (norm_sq_eq_inner x).symm end @@ -605,8 +607,7 @@ lemma inner_mul_inner_self_le (x y : E) : abs ⟪x, y⟫ * abs ⟪y, x⟫ ≤ re begin by_cases hy : y = 0, { rw [hy], simp only [is_R_or_C.abs_zero, inner_zero_left, mul_zero, add_monoid_hom.map_zero] }, - { change y ≠ 0 at hy, - have hy' : ⟪y, y⟫ ≠ 0 := λ h, by rw [inner_self_eq_zero] at h; exact hy h, + { have hy' : ⟪y, y⟫ ≠ 0 := inner_self_eq_zero.not.2 hy, set T := ⟪y, x⟫ / ⟪y, y⟫ with hT, have h₁ : re ⟪y, x⟫ = re ⟪x, y⟫ := inner_re_symm, have h₂ : im ⟪y, x⟫ = -im ⟪x, y⟫ := inner_im_symm, @@ -615,7 +616,7 @@ begin have : ⟪y, y⟫ / (⟪y, y⟫ * ⟪y, y⟫) = 1 / ⟪y, y⟫ := by rw [div_mul_eq_div_mul_one_div, div_self hy', one_mul], rw [this, div_eq_mul_inv, one_mul, ←div_eq_mul_inv] }, - have h₄ : ⟪y, y⟫ = re ⟪y, y⟫ := by simp, + have h₄ : ⟪y, y⟫ = re ⟪y, y⟫ := inner_self_re_to_K.symm, have h₅ : re ⟪y, y⟫ > 0, { refine lt_of_le_of_ne inner_self_nonneg _, intro H, @@ -634,7 +635,8 @@ begin ... = re ⟪x, x⟫ - re (T† * ⟪y, x⟫) - re (T * ⟪x, y⟫) + re (T * T† * ⟪y, y⟫) : by simp only [inner_smul_left, inner_smul_right, mul_assoc] ... = re ⟪x, x⟫ - re (⟪x, y⟫ / ⟪y, y⟫ * ⟪y, x⟫) - : by field_simp [-mul_re, hT, map_div₀, h₁, h₃, inner_conj_sym] + : by simp only [map_div₀, h₃, inner_conj_sym, sub_add_cancel] + with field_simps {discharger := tactic.field_simp.ne_zero} ... = re ⟪x, x⟫ - re (⟪x, y⟫ * ⟪y, x⟫ / ⟪y, y⟫) : by rw ←mul_div_right_comm ... = re ⟪x, x⟫ - re (⟪x, y⟫ * ⟪y, x⟫ / re ⟪y, y⟫) @@ -766,7 +768,8 @@ vectors picks out the coefficient of that vector. -/ lemma orthonormal.inner_left_sum {v : ι → E} (hv : orthonormal 𝕜 v) (l : ι → 𝕜) {s : finset ι} {i : ι} (hi : i ∈ s) : ⟪∑ i in s, (l i) • (v i), v i⟫ = conj (l i) := -by classical; simp [sum_inner, inner_smul_left, orthonormal_iff_ite.mp hv, hi] +by classical; simp only +[sum_inner, inner_smul_left, orthonormal_iff_ite.mp hv, hi, mul_boole, finset.sum_ite_eq', if_true] /-- The inner product of a linear combination of a set of orthonormal vectors with one of those vectors picks out the coefficient of that vector. -/ @@ -780,14 +783,14 @@ a sum over the first `finsupp`. -/ lemma orthonormal.inner_finsupp_eq_sum_left {v : ι → E} (hv : orthonormal 𝕜 v) (l₁ l₂ : ι →₀ 𝕜) : ⟪finsupp.total ι E 𝕜 v l₁, finsupp.total ι E 𝕜 v l₂⟫ = l₁.sum (λ i y, conj y * l₂ i) := -by simp [finsupp.total_apply _ l₁, finsupp.sum_inner, hv.inner_right_finsupp] +by simp only [l₁.total_apply _, finsupp.sum_inner, hv.inner_right_finsupp, smul_eq_mul] /-- The inner product of two linear combinations of a set of orthonormal vectors, expressed as a sum over the second `finsupp`. -/ lemma orthonormal.inner_finsupp_eq_sum_right {v : ι → E} (hv : orthonormal 𝕜 v) (l₁ l₂ : ι →₀ 𝕜) : ⟪finsupp.total ι E 𝕜 v l₁, finsupp.total ι E 𝕜 v l₂⟫ = l₂.sum (λ i y, conj (l₁ i) * y) := -by simp [finsupp.total_apply _ l₂, finsupp.inner_sum, hv.inner_left_finsupp, mul_comm] +by simp only [l₂.total_apply _, finsupp.inner_sum, hv.inner_left_finsupp, mul_comm, smul_eq_mul] /-- The inner product of two linear combinations of a set of orthonormal vectors, expressed as a sum. -/ @@ -816,7 +819,7 @@ begin intros l hl, ext i, have key : ⟪v i, finsupp.total ι E 𝕜 v l⟫ = ⟪v i, 0⟫ := by rw hl, - simpa [hv.inner_right_finsupp] using key + simpa only [hv.inner_right_finsupp, inner_zero_right] using key end /-- A subfamily of an orthonormal family (i.e., a composition with an injective map) is an @@ -857,7 +860,7 @@ lemma orthonormal.inner_finsupp_eq_zero ⟪finsupp.total ι E 𝕜 v l, v i⟫ = 0 := begin rw finsupp.mem_supported' at hl, - simp [hv.inner_left_finsupp, hl i hi], + simp only [hv.inner_left_finsupp, hl i hi, map_zero], end /-- Given an orthonormal family, a second family of vectors is orthonormal if every vector equals @@ -869,7 +872,8 @@ begin rw orthonormal_iff_ite at *, intros i j, cases hw i with hi hi; cases hw j with hj hj; split_ifs with h; - simpa [hi, hj, h] using hv i j + simpa only [hi, hj, h, inner_neg_right, inner_neg_left, + neg_neg, eq_self_iff_true, neg_eq_zero] using hv i j end /- The material that follows, culminating in the existence of a maximal orthonormal subset, is @@ -942,7 +946,7 @@ lemma norm_eq_sqrt_inner (x : E) : ‖x‖ = sqrt (re ⟪x, x⟫) := begin have h₁ : ‖x‖^2 = re ⟪x, x⟫ := norm_sq_eq_inner x, have h₂ := congr_arg sqrt h₁, - simpa using h₂, + simpa only [sqrt_sq, norm_nonneg] using h₂, end lemma norm_eq_sqrt_real_inner (x : F) : ‖x‖ = sqrt ⟪x, x⟫_ℝ := @@ -993,7 +997,7 @@ begin rw [inner_sub_sub_self], calc re (⟪x, x⟫ - ⟪x, y⟫ - ⟪y, x⟫ + ⟪y, y⟫) - = re ⟪x, x⟫ - re ⟪x, y⟫ - re ⟪y, x⟫ + re ⟪y, y⟫ : by simp + = re ⟪x, x⟫ - re ⟪x, y⟫ - re ⟪y, x⟫ + re ⟪y, y⟫ : by simp only [map_add, map_sub] ... = -re ⟪y, x⟫ - re ⟪x, y⟫ + re ⟪x, x⟫ + re ⟪y, y⟫ : by ring ... = -re (⟪x, y⟫†) - re ⟪x, y⟫ + re ⟪x, x⟫ + re ⟪y, y⟫ : by rw [inner_conj_sym] ... = -re ⟪x, y⟫ - re ⟪x, y⟫ + re ⟪x, x⟫ + re ⟪y, y⟫ : by rw [conj_re] @@ -1300,16 +1304,19 @@ basis.equiv_apply _ _ _ _ @[simp] lemma orthonormal.equiv_refl {v : basis ι 𝕜 E} (hv : orthonormal 𝕜 v) : hv.equiv hv (equiv.refl ι) = linear_isometry_equiv.refl 𝕜 E := -v.ext_linear_isometry_equiv $ λ i, by simp +v.ext_linear_isometry_equiv $ λ i, + by simp only [orthonormal.equiv_apply, equiv.coe_refl, id.def, linear_isometry_equiv.coe_refl] @[simp] lemma orthonormal.equiv_symm {v : basis ι 𝕜 E} (hv : orthonormal 𝕜 v) {v' : basis ι' 𝕜 E'} (hv' : orthonormal 𝕜 v') (e : ι ≃ ι') : (hv.equiv hv' e).symm = hv'.equiv hv e.symm := -v'.ext_linear_isometry_equiv $ λ i, (hv.equiv hv' e).injective (by simp) +v'.ext_linear_isometry_equiv $ λ i, (hv.equiv hv' e).injective $ + by simp only [linear_isometry_equiv.apply_symm_apply, orthonormal.equiv_apply, e.apply_symm_apply] @[simp] lemma orthonormal.equiv_trans {v : basis ι 𝕜 E} (hv : orthonormal 𝕜 v) {v' : basis ι' 𝕜 E'} (hv' : orthonormal 𝕜 v') (e : ι ≃ ι') {v'' : basis ι'' 𝕜 E''} (hv'' : orthonormal 𝕜 v'') (e' : ι' ≃ ι'') : (hv.equiv hv' e).trans (hv'.equiv hv'' e') = hv.equiv hv'' (e.trans e') := -v.ext_linear_isometry_equiv $ λ i, by simp +v.ext_linear_isometry_equiv $ λ i, + by simp only [linear_isometry_equiv.trans_apply, orthonormal.equiv_apply, e.coe_trans] lemma orthonormal.map_equiv {v : basis ι 𝕜 E} (hv : orthonormal 𝕜 v) {v' : basis ι' 𝕜 E'} (hv' : orthonormal 𝕜 v') (e : ι ≃ ι') : @@ -1400,8 +1407,8 @@ end lemma norm_sub_eq_norm_add {v w : E} (h : ⟪v, w⟫ = 0) : ‖w - v‖ = ‖w + v‖ := begin rw ←mul_self_inj_of_nonneg (norm_nonneg _) (norm_nonneg _), - simp [h, ←inner_self_eq_norm_mul_norm, inner_add_left, inner_add_right, inner_sub_left, - inner_sub_right, inner_re_symm] + simp only [h, ←inner_self_eq_norm_mul_norm, sub_neg_eq_add, sub_zero, map_sub, zero_re', zero_sub, + add_zero, map_add, inner_add_right, inner_sub_left, inner_sub_right, inner_re_symm, zero_add] end /-- The real inner product of two vectors, divided by the product of their @@ -1545,21 +1552,17 @@ Compare `inner_eq_norm_mul_iff`, which takes the stronger hypothesis `⟪x, y⟫ lemma abs_inner_eq_norm_iff (x y : E) (hx0 : x ≠ 0) (hy0 : y ≠ 0): abs ⟪x, y⟫ = ‖x‖ * ‖y‖ ↔ ∃ (r : 𝕜), r ≠ 0 ∧ y = r • x := begin - have hx0' : ‖x‖ ≠ 0 := by simp [norm_eq_zero, hx0], - have hy0' : ‖y‖ ≠ 0 := by simp [norm_eq_zero, hy0], - have hxy0 : ‖x‖ * ‖y‖ ≠ 0 := by simp [hx0', hy0'], + have hxy0 : ‖x‖ * ‖y‖ ≠ 0 := mul_ne_zero (norm_eq_zero.not.2 hx0) (norm_eq_zero.not.2 hy0), have h₁ : abs ⟪x, y⟫ = ‖x‖ * ‖y‖ ↔ abs (⟪x, y⟫ / (‖x‖ * ‖y‖)) = 1, - { refine ⟨_ ,_⟩, - { intro h, - norm_cast, + { split; intro h, + { norm_cast, rw [is_R_or_C.abs_div, h, abs_of_real, _root_.abs_mul, abs_norm_eq_norm, abs_norm_eq_norm], exact div_self hxy0 }, - { intro h, - norm_cast at h, + { norm_cast at h, rwa [is_R_or_C.abs_div, abs_of_real, _root_.abs_mul, abs_norm_eq_norm, abs_norm_eq_norm, div_eq_one_iff_eq hxy0] at h } }, rw [h₁, abs_inner_div_norm_mul_norm_eq_one_iff x y], - simp [hx0] + exact and_iff_right hx0, end /-- The inner product of two vectors, divided by the product of their @@ -1829,9 +1832,9 @@ instance is_R_or_C.inner_product_space : inner_product_space 𝕜 𝕜 := inner := λ x y, conj x * y, norm_sq_eq_inner := λ x, by { unfold inner, rw [mul_comm, mul_conj, of_real_re, norm_sq_eq_def'] }, - conj_sym := λ x y, by simp [mul_comm], - add_left := λ x y z, by simp [inner, add_mul], - smul_left := λ x y z, by simp [inner, mul_assoc] } + conj_sym := λ x y, by simp only [mul_comm, map_mul, star_ring_end_self_apply], + add_left := λ x y z, by simp only [add_mul, map_add], + smul_left := λ x y z, by simp only [mul_assoc, smul_eq_mul, map_mul] } @[simp] lemma is_R_or_C.inner_apply (x y : 𝕜) : ⟪x, y⟫ = (conj x) * y := rfl @@ -1911,8 +1914,8 @@ begin simp only [dfinsupp.sum, submodule.coe_inner, finset.sum_ite_eq, ite_eq_left_iff, dfinsupp.mem_support_to_fun], split_ifs with h h, - { simp }, - { simp [of_not_not h] }, + { simp only [linear_isometry.inner_map_map] }, + { simp only [of_not_not h, inner_zero_right] }, end omit dec_ι dec_V @@ -1923,26 +1926,27 @@ calc ⟪V i v, ∑ j : ι, V j (l j)⟫ = ∑ j : ι, ⟪V i v, V j (l j)⟫: by rw inner_sum ... = ∑ j, ite (i = j) ⟪V i v, V j (l j)⟫ 0 : congr_arg (finset.sum finset.univ) $ funext $ λ j, (hV.eq_ite v (l j)) -... = ⟪v, l i⟫ : by simp +... = ⟪v, l i⟫ : by simp only [finset.sum_ite_eq, finset.mem_univ, (V i).inner_map_map, if_true] lemma orthogonal_family.inner_sum (l₁ l₂ : Π i, G i) (s : finset ι) : ⟪∑ i in s, V i (l₁ i), ∑ j in s, V j (l₂ j)⟫ = ∑ i in s, ⟪l₁ i, l₂ i⟫ := by classical; calc ⟪∑ i in s, V i (l₁ i), ∑ j in s, V j (l₂ j)⟫ - = ∑ j in s, ∑ i in s, ⟪V i (l₁ i), V j (l₂ j)⟫ : by simp [sum_inner, inner_sum] + = ∑ j in s, ∑ i in s, ⟪V i (l₁ i), V j (l₂ j)⟫ : by simp only [sum_inner, inner_sum] ... = ∑ j in s, ∑ i in s, ite (i = j) ⟪V i (l₁ i), V j (l₂ j)⟫ 0 : begin congr' with i, congr' with j, apply hV.eq_ite, end -... = ∑ i in s, ⟪l₁ i, l₂ i⟫ : by simp [finset.sum_ite_of_true] +... = ∑ i in s, ⟪l₁ i, l₂ i⟫ : by simp only [finset.sum_ite_of_true, + finset.sum_ite_eq', linear_isometry.inner_map_map, imp_self, implies_true_iff] lemma orthogonal_family.norm_sum (l : Π i, G i) (s : finset ι) : ‖∑ i in s, V i (l i)‖ ^ 2 = ∑ i in s, ‖l i‖ ^ 2 := begin have : (‖∑ i in s, V i (l i)‖ ^ 2 : 𝕜) = ∑ i in s, ‖l i‖ ^ 2, - { simp [← inner_self_eq_norm_sq_to_K, hV.inner_sum] }, + { simp only [← inner_self_eq_norm_sq_to_K, hV.inner_sum] }, exact_mod_cast this, end @@ -1958,12 +1962,12 @@ lemma orthogonal_family.orthonormal_sigma_orthonormal {α : ι → Type*} {v_fam begin split, { rintros ⟨i, v⟩, - simpa using (hv_family i).1 v }, + simpa only [linear_isometry.norm_map] using (hv_family i).left v }, rintros ⟨i, v⟩ ⟨j, w⟩ hvw, by_cases hij : i = j, { subst hij, - have : v ≠ w := by simpa using hvw, - simpa using (hv_family i).2 this }, + have : v ≠ w := λ h, by { subst h, exact hvw rfl }, + simpa only [linear_isometry.inner_map_map] using (hv_family i).2 this }, { exact hV hij (v_family i v) (v_family j w) } end @@ -1978,20 +1982,20 @@ begin have hF₂ : ∀ i ∈ s₂ \ s₁, F i = - f i := λ i hi, if_neg (finset.mem_sdiff.mp hi).2, have hF : ∀ i, ‖F i‖ = ‖f i‖, { intros i, - dsimp [F], + dsimp only [F], split_ifs; - simp, }, + simp only [eq_self_iff_true, norm_neg], }, have : ‖∑ i in s₁ \ s₂, V i (F i) + ∑ i in s₂ \ s₁, V i (F i)‖ ^ 2 = ∑ i in s₁ \ s₂, ‖F i‖ ^ 2 + ∑ i in s₂ \ s₁, ‖F i‖ ^ 2, { have hs : disjoint (s₁ \ s₂) (s₂ \ s₁) := disjoint_sdiff_sdiff, simpa only [finset.sum_union hs] using hV.norm_sum F (s₁ \ s₂ ∪ s₂ \ s₁) }, convert this using 4, { refine finset.sum_congr rfl (λ i hi, _), - simp [hF₁ i hi] }, + simp only [hF₁ i hi] }, { refine finset.sum_congr rfl (λ i hi, _), - simp [hF₂ i hi] }, - { simp [hF] }, - { simp [hF] }, + simp only [hF₂ i hi, linear_isometry.map_neg] }, + { simp only [hF] }, + { simp only [hF] }, end omit dec_ι @@ -2059,11 +2063,11 @@ begin rw linear_map.mem_ker at hv, ext i, suffices : ⟪(v i : E), v i⟫ = 0, - { simpa using this }, + { simpa only [inner_self_eq_zero] using this }, calc ⟪(v i : E), v i⟫ = ⟪(v i : E), dfinsupp.lsum ℕ (λ i, (V i).subtype) v⟫ : - by simpa [dfinsupp.sum_add_hom_apply, submodule.coe_subtype] + by simpa only [dfinsupp.sum_add_hom_apply, dfinsupp.lsum_apply_apply] using (hV.inner_right_dfinsupp v i (v i)).symm - ... = 0 : by simp [hv], + ... = 0 : by simp only [hv, inner_zero_right], end include dec_ι @@ -2073,7 +2077,7 @@ lemma direct_sum.is_internal.collected_basis_orthonormal {V : ι → submodule {α : ι → Type*} {v_family : Π i, basis (α i) 𝕜 (V i)} (hv_family : ∀ i, orthonormal 𝕜 (v_family i)) : orthonormal 𝕜 (hV_sum.collected_basis v_family) := -by simpa using hV.orthonormal_sigma_orthonormal hv_family +by simpa only [hV_sum.collected_basis_coe] using hV.orthonormal_sigma_orthonormal hv_family end orthogonal_family @@ -2099,10 +2103,10 @@ def inner_product_space.is_R_or_C_to_real : inner_product_space ℝ E := conj_sym := λ x y, inner_re_symm, add_left := λ x y z, by { change re ⟪x + y, z⟫ = re ⟪x, z⟫ + re ⟪y, z⟫, - simp [inner_add_left] }, + simp only [inner_add_left, map_add] }, smul_left := λ x y r, by { change re ⟪(r : 𝕜) • x, y⟫ = r * re ⟪x, y⟫, - simp [inner_smul_left] }, + simp only [inner_smul_left, conj_of_real, of_real_mul_re] }, ..has_inner.is_R_or_C_to_real 𝕜 E, ..normed_space.restrict_scalars ℝ 𝕜 E } From 46a64b5b4268c594af770c44d9e502afc6a515cb Mon Sep 17 00:00:00 2001 From: Gabriel Ebner Date: Tue, 27 Dec 2022 16:45:56 +0000 Subject: [PATCH 0113/1291] feat(*): bump to lean 3.50.3 (#18016) Co-authored-by: Junyan Xu --- leanpkg.toml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/leanpkg.toml b/leanpkg.toml index ac4a167bb424a..98c2e9e127eb3 100644 --- a/leanpkg.toml +++ b/leanpkg.toml @@ -1,7 +1,7 @@ [package] name = "mathlib" version = "0.1" -lean_version = "leanprover-community/lean:3.50.2" +lean_version = "leanprover-community/lean:3.50.3" path = "src" [dependencies] From 422e70f7ce183d2900c586a8cda8381e788a0c62 Mon Sep 17 00:00:00 2001 From: Oliver Nash Date: Wed, 28 Dec 2022 13:01:40 +0000 Subject: [PATCH 0114/1291] feat(topology/instances/add_circle): expand API for finite-order points on the circle (#17986) --- src/data/set/finite.lean | 11 ++++ src/data/set/intervals/infinite.lean | 2 +- src/group_theory/order_of_element.lean | 54 +++++++++++++++++--- src/topology/instances/add_circle.lean | 71 ++++++++++++++++++++++++++ 4 files changed, 130 insertions(+), 8 deletions(-) diff --git a/src/data/set/finite.lean b/src/data/set/finite.lean index 33841e128ac1c..682155ac5babf 100644 --- a/src/data/set/finite.lean +++ b/src/data/set/finite.lean @@ -945,6 +945,17 @@ lemma infinite.exists_not_mem_finset {s : set α} (hs : s.infinite) (f : finset let ⟨a, has, haf⟩ := (hs.diff (to_finite f)).nonempty in ⟨a, has, λ h, haf $ finset.mem_coe.1 h⟩ +lemma not_inj_on_infinite_finite_image {f : α → β} {s : set α} + (h_inf : s.infinite) (h_fin : (f '' s).finite) : + ¬ inj_on f s := +begin + haveI : finite (f '' s) := finite_coe_iff.mpr h_fin, + haveI : infinite s := infinite_coe_iff.mpr h_inf, + have := not_injective_infinite_finite + ((f '' s).cod_restrict (s.restrict f) $ λ x, ⟨x, x.property, rfl⟩), + contrapose! this, + rwa [injective_cod_restrict, ← inj_on_iff_injective], +end /-! ### Order properties -/ diff --git a/src/data/set/intervals/infinite.lean b/src/data/set/intervals/infinite.lean index 87eef9ebec8fe..d86ec7f436560 100644 --- a/src/data/set/intervals/infinite.lean +++ b/src/data/set/intervals/infinite.lean @@ -52,7 +52,7 @@ instance [no_min_order α] {a : α} : infinite (Iic a) := no_min_order.infinite lemma Iic_infinite [no_min_order α] (a : α) : (Iic a).infinite := infinite_coe_iff.1 Iic.infinite instance [no_max_order α] {a : α} : infinite (Ioi a) := no_max_order.infinite -lemma Ioi_infinite [no_min_order α] (a : α) : (Iio a).infinite := infinite_coe_iff.1 Iio.infinite +lemma Ioi_infinite [no_max_order α] (a : α) : (Ioi a).infinite := infinite_coe_iff.1 Ioi.infinite instance [no_max_order α] {a : α} : infinite (Ici a) := no_max_order.infinite lemma Ici_infinite [no_max_order α] (a : α) : (Ici a).infinite := infinite_coe_iff.1 Ici.infinite diff --git a/src/group_theory/order_of_element.lean b/src/group_theory/order_of_element.lean index 13773d1c59379..73cb675dc908a 100644 --- a/src/group_theory/order_of_element.lean +++ b/src/group_theory/order_of_element.lean @@ -6,6 +6,7 @@ Authors: Johannes Hölzl, Julian Kuelshammer import algebra.hom.iterate import data.nat.modeq import data.set.pointwise.basic +import data.set.intervals.infinite import dynamics.periodic_pts import group_theory.index @@ -68,6 +69,19 @@ lemma is_of_fin_order_iff_pow_eq_one (x : G) : is_of_fin_order x ↔ ∃ n, 0 < n ∧ x ^ n = 1 := by { convert iff.rfl, simp [is_periodic_pt_mul_iff_pow_eq_one] } +/-- See also `injective_pow_iff_not_is_of_fin_order`. -/ +@[to_additive not_is_of_fin_add_order_of_injective_nsmul "See also +`injective_nsmul_iff_not_is_of_fin_add_order`."] +lemma not_is_of_fin_order_of_injective_pow {x : G} (h : injective (λ (n : ℕ), x^n)) : + ¬ is_of_fin_order x := +begin + simp_rw [is_of_fin_order_iff_pow_eq_one, not_exists, not_and], + intros n hn_pos hnx, + rw ← pow_zero x at hnx, + rw h hnx at hn_pos, + exact irrefl 0 hn_pos, +end + /-- Elements of finite order are of finite order in submonoids.-/ @[to_additive is_of_fin_add_order_iff_coe "Elements of finite order are of finite order in submonoids."] @@ -380,9 +394,11 @@ lemma mem_powers_iff_mem_range_order_of' [decidable_eq G] (hx : 0 < order_of x) y ∈ submonoid.powers x ↔ y ∈ (finset.range (order_of x)).image ((^) x : ℕ → G) := finset.mem_range_iff_mem_finset_range_of_mod_eq' hx (λ i, pow_eq_mod_order_of.symm) +@[to_additive] lemma pow_eq_one_iff_modeq : x ^ n = 1 ↔ n ≡ 0 [MOD (order_of x)] := by rw [modeq_zero_iff_dvd, order_of_dvd_iff_pow_eq_one] +@[to_additive] lemma pow_eq_pow_iff_modeq : x ^ n = x ^ m ↔ n ≡ m [MOD (order_of x)] := begin wlog hmn : m ≤ n, @@ -391,6 +407,34 @@ begin exact ⟨λ h, nat.modeq.add_left _ h, λ h, nat.modeq.add_left_cancel' _ h⟩, end +@[simp, to_additive injective_nsmul_iff_not_is_of_fin_add_order] +lemma injective_pow_iff_not_is_of_fin_order {x : G} : + injective (λ (n : ℕ), x^n) ↔ ¬ is_of_fin_order x := +begin + refine ⟨λ h, not_is_of_fin_order_of_injective_pow h, λ h n m hnm, _⟩, + rwa [pow_eq_pow_iff_modeq, order_of_eq_zero_iff.mpr h, modeq_zero_iff] at hnm, +end + +@[to_additive infinite_not_is_of_fin_add_order] +lemma infinite_not_is_of_fin_order {x : G} (h : ¬ is_of_fin_order x) : + {y : G | ¬ is_of_fin_order y}.infinite := +begin + let s := {n | 0 < n}.image (λ (n : ℕ), x^n), + have hs : s ⊆ {y : G | ¬ is_of_fin_order y}, + { rintros - ⟨n, hn : 0 < n, rfl⟩ (contra : is_of_fin_order (x^n)), + apply h, + rw is_of_fin_order_iff_pow_eq_one at contra ⊢, + obtain ⟨m, hm, hm'⟩ := contra, + exact ⟨n * m, mul_pos hn hm, by rwa pow_mul⟩, }, + suffices : s.infinite, { exact this.mono hs, }, + contrapose! h, + have : ¬ injective (λ (n : ℕ), x^n), + { have := set.not_inj_on_infinite_finite_image (set.Ioi_infinite 0) (set.not_infinite.mp h), + contrapose! this, + exact set.inj_on_of_injective this _, }, + rwa [injective_pow_iff_not_is_of_fin_order, not_not] at this, +end + end cancel_monoid section group @@ -541,13 +585,9 @@ variables [left_cancel_monoid G] [add_left_cancel_monoid A] @[to_additive] lemma exists_pow_eq_one [finite G] (x : G) : is_of_fin_order x := begin - refine (is_of_fin_order_iff_pow_eq_one _).mpr _, - obtain ⟨i, j, a_eq, ne⟩ : ∃(i j : ℕ), x ^ i = x ^ j ∧ i ≠ j := - by simpa only [not_forall, exists_prop, injective] - using (not_injective_infinite_finite (λi:ℕ, x^i)), - wlog h'' : j ≤ i, - refine ⟨i - j, tsub_pos_of_lt (lt_of_le_of_ne h'' ne.symm), mul_right_injective (x^j) _⟩, - rw [mul_one, ← pow_add, ← a_eq, add_tsub_cancel_of_le h''], + have : (set.univ : set G).finite := set.univ.to_finite, + contrapose! this, + exact set.infinite.mono (set.subset_univ _) (infinite_not_is_of_fin_order this), end @[to_additive add_order_of_le_card_univ] diff --git a/src/topology/instances/add_circle.lean b/src/topology/instances/add_circle.lean index a7f8b9a387576..caeacf8833403 100644 --- a/src/topology/instances/add_circle.lean +++ b/src/topology/instances/add_circle.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 Oliver Nash. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Oliver Nash -/ +import data.nat.totient import algebra.ring.add_aut import group_theory.divisible import group_theory.order_of_element @@ -67,6 +68,10 @@ lemma coe_nsmul {n : ℕ} {x : 𝕜} : (↑(n • x) : add_circle p) = n • (x lemma coe_zsmul {n : ℤ} {x : 𝕜} : (↑(n • x) : add_circle p) = n • (x : add_circle p) := rfl +lemma coe_add (x y : 𝕜) : (↑(x + y) : add_circle p) = (x : add_circle p) + (y : add_circle p) := rfl + +lemma coe_sub (x y : 𝕜) : (↑(x - y) : add_circle p) = (x : add_circle p) - (y : add_circle p) := rfl + lemma coe_neg {x : 𝕜} : (↑(-x) : add_circle p) = -(x : add_circle p) := rfl lemma coe_eq_zero_iff {x : 𝕜} : (x : add_circle p) = 0 ↔ ∃ (n : ℤ), n • p = x := @@ -280,6 +285,72 @@ begin simpa [zero_add] using (equiv_Ico p 0 u).2.2, }, end +lemma add_order_of_eq_pos_iff {u : add_circle p} {n : ℕ} (h : 0 < n) : + add_order_of u = n ↔ ∃ m < n, gcd m n = 1 ∧ ↑(↑m / ↑n * p) = u := +begin + refine ⟨λ hu, _, _⟩, + { rw ← hu at h, + obtain ⟨m, h₀, h₁, h₂⟩ := exists_gcd_eq_one_of_is_of_fin_add_order (add_order_of_pos_iff.mp h), + refine ⟨m, _, _, _⟩; + rwa ← hu, }, + { rintros ⟨m, h₀, h₁, rfl⟩, + exact add_order_of_div_of_gcd_eq_one h h₁, }, +end + +variables (p) + +/-- The natural bijection between points of order `n` and natural numbers less than and coprime to +`n`. The inverse of the map sends `m ↦ (m/n * p : add_circle p)` where `m` is coprime to `n` and +satisfies `0 ≤ m < n`. -/ +def set_add_order_of_equiv {n : ℕ} (hn : 0 < n) : + {u : add_circle p | add_order_of u = n} ≃ {m | m < n ∧ gcd m n = 1} := +{ to_fun := λ u, by + { let h := (add_order_of_eq_pos_iff hn).mp u.property, + exact ⟨classical.some h, classical.some (classical.some_spec h), + (classical.some_spec (classical.some_spec h)).1⟩, }, + inv_fun := λ m, ⟨↑((m : 𝕜) / n * p), add_order_of_div_of_gcd_eq_one hn (m.property.2)⟩, + left_inv := λ u, subtype.ext + (classical.some_spec (classical.some_spec $ (add_order_of_eq_pos_iff hn).mp u.2)).2, + right_inv := + begin + rintros ⟨m, hm₁, hm₂⟩, + let u : {u : add_circle p | add_order_of u = n} := + ⟨↑((m : 𝕜) / n * p), add_order_of_div_of_gcd_eq_one hn hm₂⟩, + let h := (add_order_of_eq_pos_iff hn).mp u.property, + ext, + let m' := classical.some h, + change m' = m, + obtain ⟨h₁ : m' < n, h₂ : gcd m' n = 1, h₃ : quotient_add_group.mk ((m' : 𝕜) / n * p) = + quotient_add_group.mk ((m : 𝕜) / n * p)⟩ := classical.some_spec h, + replace h₃ := congr_arg (coe : Ico 0 (0 + p) → 𝕜) (congr_arg (equiv_Ico p 0) h₃), + simpa only [coe_equiv_Ico_mk_apply, mul_left_inj' hp.out.ne', mul_div_cancel _ hp.out.ne', + int.fract_div_nat_cast_eq_div_nat_cast_mod, + div_left_inj' (nat.cast_ne_zero.mpr hn.ne' : (n : 𝕜) ≠ 0), nat.cast_inj, + (nat.mod_eq_iff_lt hn.ne').mpr hm₁, (nat.mod_eq_iff_lt hn.ne').mpr h₁] using h₃, + end } + +@[simp] lemma card_add_order_of_eq_totient {n : ℕ} : + nat.card {u : add_circle p // add_order_of u = n} = n.totient := +begin + rcases n.eq_zero_or_pos with rfl | hn, + { simp only [nat.totient_zero, add_order_of_eq_zero_iff], + rcases em (∃ (u : add_circle p), ¬ is_of_fin_add_order u) with ⟨u, hu⟩ | h, + { haveI : infinite {u : add_circle p // ¬is_of_fin_add_order u}, + { erw infinite_coe_iff, + exact infinite_not_is_of_fin_add_order hu, }, + exact nat.card_eq_zero_of_infinite, }, + { haveI : is_empty {u : add_circle p // ¬is_of_fin_add_order u}, { simpa using h, }, + exact nat.card_of_is_empty, }, }, + { rw [← coe_set_of, nat.card_congr (set_add_order_of_equiv p hn), + n.totient_eq_card_lt_and_coprime], + simpa only [@nat.coprime_comm _ n], }, +end + +lemma finite_set_of_add_order_eq {n : ℕ} (hn : 0 < n) : + {u : add_circle p | add_order_of u = n}.finite := +finite_coe_iff.mp $ nat.finite_of_card_ne_zero $ by simpa only [coe_set_of, + card_add_order_of_eq_totient p] using (nat.totient_pos hn).ne' + end finite_order_points end linear_ordered_field From 986c4d5761f938b2e1c43c01f001b6d9d88c2055 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Thu, 29 Dec 2022 17:01:52 +0000 Subject: [PATCH 0115/1291] fix(*): mark some classes as Prop (#18015) --- .../100-theorems-list/30_ballot_problem.lean | 2 +- src/algebra/hom/freiman.lean | 6 +++--- .../projective_spectrum/scheme.lean | 8 ++++---- src/category_theory/monoidal/linear.lean | 4 ++-- src/category_theory/monoidal/preadditive.lean | 4 ++-- src/field_theory/subfield.lean | 2 +- src/group_theory/subgroup/basic.lean | 18 +++++------------- src/group_theory/submonoid/basic.lean | 15 ++++----------- src/group_theory/subsemigroup/basic.lean | 4 ++-- src/ring_theory/subring/basic.lean | 3 +-- src/ring_theory/subsemiring/basic.lean | 6 ++---- 11 files changed, 27 insertions(+), 45 deletions(-) diff --git a/archive/100-theorems-list/30_ballot_problem.lean b/archive/100-theorems-list/30_ballot_problem.lean index 4891ab851160e..6e9e589531cce 100644 --- a/archive/100-theorems-list/30_ballot_problem.lean +++ b/archive/100-theorems-list/30_ballot_problem.lean @@ -277,7 +277,7 @@ private def measureable_space_list_int : measurable_space (list ℤ) := ⊤ local attribute [instance] measureable_space_list_int -private def measurable_singleton_class_list_int : measurable_singleton_class (list ℤ) := +private lemma measurable_singleton_class_list_int : measurable_singleton_class (list ℤ) := { measurable_set_singleton := λ s, trivial } local attribute [instance] measurable_singleton_class_list_int diff --git a/src/algebra/hom/freiman.lean b/src/algebra/hom/freiman.lean index 3f880ffd34a70..f55b01bfd49b5 100644 --- a/src/algebra/hom/freiman.lean +++ b/src/algebra/hom/freiman.lean @@ -72,7 +72,7 @@ notation (name := freiman_hom) A ` →*[`:25 n:25 `] `:0 β:0 := freiman_hom A /-- `add_freiman_hom_class F s β n` states that `F` is a type of `n`-ary sums-preserving morphisms. You should extend this class when you extend `add_freiman_hom`. -/ class add_freiman_hom_class (F : Type*) (A : out_param $ set α) (β : out_param $ Type*) - [add_comm_monoid α] [add_comm_monoid β] (n : ℕ) [fun_like F α (λ _, β)] := + [add_comm_monoid α] [add_comm_monoid β] (n : ℕ) [fun_like F α (λ _, β)] : Prop := (map_sum_eq_map_sum' (f : F) {s t : multiset α} (hsA : ∀ ⦃x⦄, x ∈ s → x ∈ A) (htA : ∀ ⦃x⦄, x ∈ t → x ∈ A) (hs : s.card = n) (ht : t.card = n) (h : s.sum = t.sum) : (s.map f).sum = (t.map f).sum) @@ -83,7 +83,7 @@ You should extend this class when you extend `freiman_hom`. -/ "`add_freiman_hom_class F A β n` states that `F` is a type of `n`-ary sums-preserving morphisms. You should extend this class when you extend `add_freiman_hom`."] class freiman_hom_class (F : Type*) (A : out_param $ set α) (β : out_param $ Type*) [comm_monoid α] - [comm_monoid β] (n : ℕ) [fun_like F α (λ _, β)] := + [comm_monoid β] (n : ℕ) [fun_like F α (λ _, β)] : Prop := (map_prod_eq_map_prod' (f : F) {s t : multiset α} (hsA : ∀ ⦃x⦄, x ∈ s → x ∈ A) (htA : ∀ ⦃x⦄, x ∈ t → x ∈ A) (hs : s.card = n) (ht : t.card = n) (h : s.prod = t.prod) : (s.map f).prod = (t.map f).prod) @@ -369,7 +369,7 @@ def freiman_hom.to_freiman_hom (h : m ≤ n) (f : A →*[n] β) : A →*[m] β : /-- A `n`-Freiman homomorphism is also a `m`-Freiman homomorphism for any `m ≤ n`. -/ @[to_additive add_freiman_hom.add_freiman_hom_class_of_le "An additive `n`-Freiman homomorphism is also an additive `m`-Freiman homomorphism for any `m ≤ n`."] -def freiman_hom.freiman_hom_class_of_le [freiman_hom_class F A β n] (h : m ≤ n) : +lemma freiman_hom.freiman_hom_class_of_le [freiman_hom_class F A β n] (h : m ≤ n) : freiman_hom_class F A β m := { map_prod_eq_map_prod' := λ f s t hsA htA hs ht hst, map_prod_eq_map_prod_of_le f hsA htA hs ht hst h } diff --git a/src/algebraic_geometry/projective_spectrum/scheme.lean b/src/algebraic_geometry/projective_spectrum/scheme.lean index 1ae0d6c22164a..3d28c47c871a3 100644 --- a/src/algebraic_geometry/projective_spectrum/scheme.lean +++ b/src/algebraic_geometry/projective_spectrum/scheme.lean @@ -196,8 +196,8 @@ begin rw set.not_disjoint_iff_nonempty_inter, refine ⟨f^N * f^M, eq1.symm ▸ mul_mem_right _ _ (sum_mem _ (λ i hi, mul_mem_left _ _ _)), ⟨N+M, by rw pow_add⟩⟩, - generalize_proofs h, - exact (classical.some_spec h).1, + generalize_proofs h₁ h₂, + exact (classical.some_spec h₂).1, end variable (f) @@ -236,7 +236,7 @@ def to_fun (x : Proj.T| (pbo f)) : (Spec.T (A⁰_ f)) := { exact false.elim (x.2 (x.1.is_prime.mem_of_pow_mem M rid2)), }, { rw [mul_comm _ (f^N), eq1], refine mul_mem_right _ _ (mul_mem_right _ _ (sum_mem _ (λ i hi, mul_mem_left _ _ _))), - generalize_proofs h, exact (classical.some_spec h).1 }, + generalize_proofs h₁ h₂, exact (classical.some_spec h₂).1 }, end⟩ /- @@ -283,7 +283,7 @@ begin { exact y.2 (y.1.is_prime.mem_of_pow_mem M H3), }, { rw [mul_comm _ (f^N), eq1], refine mul_mem_right _ _ (mul_mem_right _ _ (sum_mem _ (λ i hi, mul_mem_left _ _ _))), - generalize_proofs h, exact (classical.some_spec h).1, }, }, + generalize_proofs h₁ h₂, exact (classical.some_spec h₂).1, }, }, end end to_Spec diff --git a/src/category_theory/monoidal/linear.lean b/src/category_theory/monoidal/linear.lean index 75d95bead92ff..adf3555e1fcec 100644 --- a/src/category_theory/monoidal/linear.lean +++ b/src/category_theory/monoidal/linear.lean @@ -25,7 +25,7 @@ variables [monoidal_category C] [monoidal_preadditive C] /-- A category is `monoidal_linear R` if tensoring is `R`-linear in both factors. -/ -class monoidal_linear := +class monoidal_linear : Prop := (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), @@ -44,7 +44,7 @@ instance tensoring_right_linear (X : C) : ((tensoring_right C).obj X).linear R : /-- A faithful linear monoidal functor to a linear monoidal category ensures that the domain is linear monoidal. -/ -def monoidal_linear_of_faithful +lemma monoidal_linear_of_faithful {D : Type*} [category D] [preadditive D] [linear R D] [monoidal_category D] [monoidal_preadditive D] (F : monoidal_functor D C) [faithful F.to_functor] diff --git a/src/category_theory/monoidal/preadditive.lean b/src/category_theory/monoidal/preadditive.lean index 10c0e82a49605..9fb5c2b68234f 100644 --- a/src/category_theory/monoidal/preadditive.lean +++ b/src/category_theory/monoidal/preadditive.lean @@ -29,7 +29,7 @@ A category is `monoidal_preadditive` if tensoring is additive in both factors. Note we don't `extend preadditive C` here, as `abelian C` already extends it, and we'll need to have both typeclasses sometimes. -/ -class monoidal_preadditive := +class monoidal_preadditive : Prop := (tensor_zero' : ∀ {W X Y Z : C} (f : W ⟶ X), f ⊗ (0 : Y ⟶ Z) = 0 . obviously) (zero_tensor' : ∀ {W X Y Z : C} (f : Y ⟶ Z), (0 : W ⟶ X) ⊗ f = 0 . obviously) (tensor_add' : ∀ {W X Y Z : C} (f : W ⟶ X) (g h : Y ⟶ Z), f ⊗ (g + h) = f ⊗ g + f ⊗ h . obviously) @@ -52,7 +52,7 @@ instance tensoring_right_additive (X : C) : ((tensoring_right C).obj X).additive /-- A faithful additive monoidal functor to a monoidal preadditive category ensures that the domain is monoidal preadditive. -/ -def monoidal_preadditive_of_faithful {D : Type*} [category D] [preadditive D] [monoidal_category D] +lemma monoidal_preadditive_of_faithful {D} [category D] [preadditive D] [monoidal_category D] (F : monoidal_functor D C) [faithful F.to_functor] [F.to_functor.additive] : monoidal_preadditive D := { tensor_zero' := by { intros, apply F.to_functor.map_injective, simp [F.map_tensor], }, diff --git a/src/field_theory/subfield.lean b/src/field_theory/subfield.lean index d704fc08fbd46..1f66d94ff24e0 100644 --- a/src/field_theory/subfield.lean +++ b/src/field_theory/subfield.lean @@ -66,7 +66,7 @@ variables {K : Type u} {L : Type v} {M : Type w} [field K] [field L] [field M] /-- `subfield_class S K` states `S` is a type of subsets `s ⊆ K` closed under field operations. -/ class subfield_class (S : Type*) (K : out_param $ Type*) [field K] [set_like S K] - extends subring_class S K, inv_mem_class S K. + extends subring_class S K, inv_mem_class S K : Prop namespace subfield_class diff --git a/src/group_theory/subgroup/basic.lean b/src/group_theory/subgroup/basic.lean index 2afc191ebafc8..2cf2efd18c74e 100644 --- a/src/group_theory/subgroup/basic.lean +++ b/src/group_theory/subgroup/basic.lean @@ -95,39 +95,31 @@ variables {A : Type*} [add_group A] section subgroup_class /-- `inv_mem_class S G` states `S` is a type of subsets `s ⊆ G` closed under inverses. -/ -class inv_mem_class (S G : Type*) [has_inv G] [set_like S G] := +class inv_mem_class (S G : Type*) [has_inv G] [set_like S G] : Prop := (inv_mem : ∀ {s : S} {x}, x ∈ s → x⁻¹ ∈ s) export inv_mem_class (inv_mem) /-- `neg_mem_class S G` states `S` is a type of subsets `s ⊆ G` closed under negation. -/ -class neg_mem_class (S G : Type*) [has_neg G] [set_like S G] := +class neg_mem_class (S G : Type*) [has_neg G] [set_like S G] : Prop := (neg_mem : ∀ {s : S} {x}, x ∈ s → -x ∈ s) export neg_mem_class (neg_mem) /-- `subgroup_class S G` states `S` is a type of subsets `s ⊆ G` that are subgroups of `G`. -/ class subgroup_class (S G : Type*) [div_inv_monoid G] [set_like S G] - extends submonoid_class S G := -(inv_mem : ∀ {s : S} {x}, x ∈ s → x⁻¹ ∈ s) + extends submonoid_class S G, inv_mem_class S G : Prop /-- `add_subgroup_class S G` states `S` is a type of subsets `s ⊆ G` that are additive subgroups of `G`. -/ class add_subgroup_class (S G : Type*) [sub_neg_monoid G] [set_like S G] - extends add_submonoid_class S G := -(neg_mem : ∀ {s : S} {x}, x ∈ s → -x ∈ s) + extends add_submonoid_class S G, neg_mem_class S G : Prop attribute [to_additive] inv_mem_class subgroup_class -variables (M S : Type*) [div_inv_monoid M] [set_like S M] [hSM : subgroup_class S M] +variables {M S : Type*} [div_inv_monoid M] [set_like S M] [hSM : subgroup_class S M] {H K : S} include hSM -@[to_additive, priority 100] -- See note [lower instance priority] -instance subgroup_class.to_inv_mem_class : inv_mem_class S M := -{ .. hSM } - -variables {S M} {H K : S} - /-- A subgroup is closed under division. -/ @[to_additive "An additive subgroup is closed under subtraction."] theorem div_mem {x y : M} (hx : x ∈ H) (hy : y ∈ H) : x / y ∈ H := diff --git a/src/group_theory/submonoid/basic.lean b/src/group_theory/submonoid/basic.lean index afde4a523284c..8fa051273cbc5 100644 --- a/src/group_theory/submonoid/basic.lean +++ b/src/group_theory/submonoid/basic.lean @@ -60,13 +60,13 @@ variables [mul_one_class M] {s : set M} variables [add_zero_class A] {t : set A} /-- `one_mem_class S M` says `S` is a type of subsets `s ≤ M`, such that `1 ∈ s` for all `s`. -/ -class one_mem_class (S : Type*) (M : out_param $ Type*) [has_one M] [set_like S M] := +class one_mem_class (S : Type*) (M : out_param $ Type*) [has_one M] [set_like S M] : Prop := (one_mem : ∀ (s : S), (1 : M) ∈ s) export one_mem_class (one_mem) /-- `zero_mem_class S M` says `S` is a type of subsets `s ≤ M`, such that `0 ∈ s` for all `s`. -/ -class zero_mem_class (S : Type*) (M : out_param $ Type*) [has_zero M] [set_like S M] := +class zero_mem_class (S : Type*) (M : out_param $ Type*) [has_zero M] [set_like S M] : Prop := (zero_mem : ∀ (s : S), (0 : M) ∈ s) export zero_mem_class (zero_mem) @@ -90,8 +90,7 @@ add_decl_doc submonoid.to_subsemigroup /-- `submonoid_class S M` says `S` is a type of subsets `s ≤ M` that contain `1` and are closed under `(*)` -/ class submonoid_class (S : Type*) (M : out_param $ Type*) [mul_one_class M] [set_like S M] - extends mul_mem_class S M := -(one_mem : ∀ (s : S), (1 : M) ∈ s) + extends mul_mem_class S M, one_mem_class S M : Prop section @@ -112,16 +111,10 @@ add_decl_doc add_submonoid.to_add_subsemigroup /-- `add_submonoid_class S M` says `S` is a type of subsets `s ≤ M` that contain `0` and are closed under `(+)` -/ class add_submonoid_class (S : Type*) (M : out_param $ Type*) [add_zero_class M] [set_like S M] - extends add_mem_class S M := -(zero_mem : ∀ (s : S), (0 : M) ∈ s) + extends add_mem_class S M, zero_mem_class S M : Prop attribute [to_additive] submonoid submonoid_class -@[to_additive, priority 100] -- See note [lower instance priority] -instance submonoid_class.to_one_mem_class (S : Type*) (M : out_param $ Type*) [mul_one_class M] - [set_like S M] [h : submonoid_class S M] : one_mem_class S M := -{ ..h } - @[to_additive] lemma pow_mem {M} [monoid M] {A : Type*} [set_like A M] [submonoid_class A M] {S : A} {x : M} (hx : x ∈ S) : ∀ (n : ℕ), x ^ n ∈ S diff --git a/src/group_theory/subsemigroup/basic.lean b/src/group_theory/subsemigroup/basic.lean index 0c7c232cd6857..28d6e513792c9 100644 --- a/src/group_theory/subsemigroup/basic.lean +++ b/src/group_theory/subsemigroup/basic.lean @@ -54,13 +54,13 @@ variables [has_mul M] {s : set M} variables [has_add A] {t : set A} /-- `mul_mem_class S M` says `S` is a type of subsets `s ≤ M` that are closed under `(*)` -/ -class mul_mem_class (S : Type*) (M : out_param $ Type*) [has_mul M] [set_like S M] := +class mul_mem_class (S : Type*) (M : out_param $ Type*) [has_mul M] [set_like S M] : Prop := (mul_mem : ∀ {s : S} {a b : M}, a ∈ s → b ∈ s → a * b ∈ s) export mul_mem_class (mul_mem) /-- `add_mem_class S M` says `S` is a type of subsets `s ≤ M` that are closed under `(+)` -/ -class add_mem_class (S : Type*) (M : out_param $ Type*) [has_add M] [set_like S M] := +class add_mem_class (S : Type*) (M : out_param $ Type*) [has_add M] [set_like S M] : Prop := (add_mem : ∀ {s : S} {a b : M}, a ∈ s → b ∈ s → a + b ∈ s) export add_mem_class (add_mem) diff --git a/src/ring_theory/subring/basic.lean b/src/ring_theory/subring/basic.lean index 44a1a9fb8c92f..54ea6c2f32073 100644 --- a/src/ring_theory/subring/basic.lean +++ b/src/ring_theory/subring/basic.lean @@ -72,8 +72,7 @@ section subring_class /-- `subring_class S R` states that `S` is a type of subsets `s ⊆ R` that are both a multiplicative submonoid and an additive subgroup. -/ class subring_class (S : Type*) (R : out_param $ Type u) [ring R] [set_like S R] - extends subsemiring_class S R := -(neg_mem : ∀ {s : S} {a : R}, a ∈ s → -a ∈ s) + extends subsemiring_class S R, neg_mem_class S R : Prop @[priority 100] -- See note [lower instance priority] instance subring_class.add_subgroup_class (S : Type*) (R : out_param $ Type u) [set_like S R] diff --git a/src/ring_theory/subsemiring/basic.lean b/src/ring_theory/subsemiring/basic.lean index 3396929e837c8..f064254ef77bd 100644 --- a/src/ring_theory/subsemiring/basic.lean +++ b/src/ring_theory/subsemiring/basic.lean @@ -31,7 +31,7 @@ section add_submonoid_with_one_class and are closed under `(+)` -/ class add_submonoid_with_one_class (S : Type*) (R : out_param $ Type*) [add_monoid_with_one R] [set_like S R] - extends add_submonoid_class S R, one_mem_class S R + extends add_submonoid_class S R, one_mem_class S R : Prop variables {S R : Type*} [add_monoid_with_one R] [set_like S R] (s : S) @@ -56,9 +56,7 @@ section subsemiring_class /-- `subsemiring_class S R` states that `S` is a type of subsets `s ⊆ R` that are both a multiplicative and an additive submonoid. -/ class subsemiring_class (S : Type*) (R : out_param $ Type u) [non_assoc_semiring R] [set_like S R] - extends submonoid_class S R := -(add_mem : ∀ {s : S} {a b : R}, a ∈ s → b ∈ s → a + b ∈ s) -(zero_mem : ∀ (s : S), (0 : R) ∈ s) + extends submonoid_class S R, add_submonoid_class S R : Prop @[priority 100] -- See note [lower instance priority] instance subsemiring_class.add_submonoid_with_one_class (S : Type*) (R : out_param $ Type u) From 09597669f02422ed388036273d8848119699c22f Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Fri, 30 Dec 2022 06:17:31 +0000 Subject: [PATCH 0116/1291] chore(group_theory/quotient_group): open `function`, use `rfl` (#18014) --- src/group_theory/quotient_group.lean | 30 +++++++++++++--------------- 1 file changed, 14 insertions(+), 16 deletions(-) diff --git a/src/group_theory/quotient_group.lean b/src/group_theory/quotient_group.lean index 18c1f5b0d9250..14fc5da61645d 100644 --- a/src/group_theory/quotient_group.lean +++ b/src/group_theory/quotient_group.lean @@ -38,6 +38,7 @@ proves Noether's first and second isomorphism theorems. isomorphism theorems, quotient groups -/ +open function universes u v namespace quotient_group @@ -69,7 +70,7 @@ lemma coe_mk' : (mk' N : G → G ⧸ N) = coe := rfl lemma mk'_apply (x : G) : mk' N x = x := rfl @[to_additive] -lemma mk'_surjective : function.surjective $ mk' N := @mk_surjective _ _ N +lemma mk'_surjective : surjective $ mk' N := @mk_surjective _ _ N @[to_additive] lemma mk'_eq_mk' {x y : G} : mk' N x = mk' N y ↔ ∃ z ∈ N, x * z = y := @@ -163,19 +164,16 @@ end @[simp, to_additive] lemma map_coe (M : subgroup H) [M.normal] (f : G →* H) (h : N ≤ M.comap f) (x : G) : map N M f h ↑x = ↑(f x) := -lift_mk' _ _ x +rfl @[to_additive] lemma map_mk' (M : subgroup H) [M.normal] (f : G →* H) (h : N ≤ M.comap f) (x : G) : map N M f h (mk' _ x) = ↑(f x) := -quotient_group.lift_mk' _ _ x +rfl @[to_additive] lemma map_id_apply (h : N ≤ subgroup.comap (monoid_hom.id _) N := (subgroup.comap_id N).le) (x) : map N N (monoid_hom.id _) h x = x := -begin - refine induction_on' x (λ x, _), - simp only [map_coe, monoid_hom.id_apply] -end +induction_on' x $ λ x, rfl @[simp, to_additive] lemma map_id (h : N ≤ subgroup.comap (monoid_hom.id _) N := (subgroup.comap_id N).le) : @@ -226,19 +224,19 @@ def congr (e : G ≃* H) (he : G'.map ↑e = H') : G ⧸ G' ≃* H ⧸ H' := @[simp] lemma congr_mk (e : G ≃* H) (he : G'.map ↑e = H') (x) : congr G' H' e he (mk x) = e x := -map_mk' G' _ _ (he ▸ G'.le_comap_map e) _ +rfl lemma congr_mk' (e : G ≃* H) (he : G'.map ↑e = H') (x) : congr G' H' e he (mk' G' x) = mk' H' (e x) := -map_mk' G' _ _ (he ▸ G'.le_comap_map e) _ +rfl @[simp] lemma congr_apply (e : G ≃* H) (he : G'.map ↑e = H') (x : G) : congr G' H' e he x = mk' H' (e x) := -map_mk' G' _ _ (he ▸ G'.le_comap_map e) _ +rfl @[simp] lemma congr_refl (he : G'.map (mul_equiv.refl G : G →* G) = G' := subgroup.map_id G') : congr G' G' (mul_equiv.refl G) he = mul_equiv.refl (G ⧸ G') := -by ext x; refine induction_on' x (λ x', _); simp +by { ext ⟨x⟩, refl } @[simp] lemma congr_symm (e : G ≃* H) (he : G'.map ↑e = H') : (congr G' H' e he).symm = congr H' G' e.symm ((subgroup.map_symm_eq_iff_map_eq _).mpr he) := @@ -248,7 +246,7 @@ end congr variables (φ : G →* H) -open function monoid_hom +open monoid_hom /-- The induced map from the quotient by the kernel to the codomain. -/ @[to_additive "The induced map from the quotient by the kernel to the codomain."] @@ -301,11 +299,11 @@ mul_equiv.of_bijective (range_ker_lift φ) ⟨range_ker_lift_injective φ, range with a right inverse `ψ : H → G`. -/ @[to_additive "The canonical isomorphism `G/(ker φ) ≃+ H` induced by a homomorphism `φ : G →+ H` with a right inverse `ψ : H → G`.", simps] -def quotient_ker_equiv_of_right_inverse (ψ : H → G) (hφ : function.right_inverse ψ φ) : +def quotient_ker_equiv_of_right_inverse (ψ : H → G) (hφ : right_inverse ψ φ) : G ⧸ ker φ ≃* H := { to_fun := ker_lift φ, inv_fun := mk ∘ ψ, - left_inv := λ x, ker_lift_injective φ (by rw [function.comp_app, ker_lift_mk', hφ]), + left_inv := λ x, ker_lift_injective φ (by rw [comp_app, ker_lift_mk', hφ]), right_inv := hφ, .. ker_lift φ } @@ -321,7 +319,7 @@ For a `computable` version, see `quotient_group.quotient_ker_equiv_of_right_inve @[to_additive "The canonical isomorphism `G/(ker φ) ≃+ H` induced by a surjection `φ : G →+ H`. For a `computable` version, see `quotient_add_group.quotient_ker_equiv_of_right_inverse`."] -noncomputable def quotient_ker_equiv_of_surjective (hφ : function.surjective φ) : +noncomputable def quotient_ker_equiv_of_surjective (hφ : surjective φ) : G ⧸ (ker φ) ≃* H := quotient_ker_equiv_of_right_inverse φ _ hφ.has_right_inverse.some_spec @@ -445,7 +443,7 @@ noncomputable def quotient_inf_equiv_prod_normal_quotient (H N : subgroup G) [N. /- φ is the natural homomorphism H →* (HN)/N. -/ let φ : H →* _ ⧸ (N.subgroup_of (H ⊔ N)) := (mk' $ N.subgroup_of (H ⊔ N)).comp (inclusion le_sup_left) in -have φ_surjective : function.surjective φ := λ x, x.induction_on' $ +have φ_surjective : surjective φ := λ x, x.induction_on' $ begin rintro ⟨y, (hy : y ∈ ↑(H ⊔ N))⟩, rw mul_normal H N at hy, rcases hy with ⟨h, n, hh, hn, rfl⟩, From a437a2499163d85d670479f69f625f461cc5fef9 Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Fri, 30 Dec 2022 09:49:10 +0000 Subject: [PATCH 0117/1291] chore(topology/instances/add_circle): golf (#18020) Net -23 lines. Also: + Switch the order of `exists_gcd_eq_one_of_is_of_fin_add_order` and `add_order_of_eq_pos_iff`. + Switch the order of `gcd_mul_add_order_of_div_eq` and `add_order_of_div_of_gcd_eq_one`. + Add easy lemmas `order_of_eq_iff`, `coe_period` and `add_order_of_period_div`. + Remove `[floor_ring]` hypothesis. + Replace `gcd` by `nat.gcd`. Co-authored-by: Oliver Nash Co-authored-by: Junyan Xu --- src/group_theory/order_of_element.lean | 9 ++ src/topology/instances/add_circle.lean | 182 ++++++++++--------------- 2 files changed, 84 insertions(+), 107 deletions(-) diff --git a/src/group_theory/order_of_element.lean b/src/group_theory/order_of_element.lean index 73cb675dc908a..c8fd227ddc3ed 100644 --- a/src/group_theory/order_of_element.lean +++ b/src/group_theory/order_of_element.lean @@ -153,6 +153,15 @@ by rwa [order_of, minimal_period, dif_neg] order_of x = 0 ↔ ∀ n : ℕ, 0 < n → x ^ n ≠ 1 := by simp_rw [order_of_eq_zero_iff, is_of_fin_order_iff_pow_eq_one, not_exists, not_and] +@[to_additive add_order_of_eq_iff] lemma order_of_eq_iff {n} (h : 0 < n) : + order_of x = n ↔ x ^ n = 1 ∧ ∀ m, m < n → 0 < m → x ^ m ≠ 1 := +begin + simp_rw [ne, ← is_periodic_pt_mul_iff_pow_eq_one, order_of, minimal_period], + split_ifs with h1, + { rw [find_eq_iff, exists_prop_of_true h], push_neg, refl }, + { rw iff_false_left h.ne, rintro ⟨h', -⟩, exact h1 ⟨n, h, h'⟩ }, +end + /-- A group element has finite order iff its order is positive. -/ @[to_additive add_order_of_pos_iff "A group element has finite additive order iff its order is positive."] diff --git a/src/topology/instances/add_circle.lean b/src/topology/instances/add_circle.lean index caeacf8833403..c9b6076421899 100644 --- a/src/topology/instances/add_circle.lean +++ b/src/topology/instances/add_circle.lean @@ -90,11 +90,11 @@ begin { exact ⟨(n : ℤ), by simp⟩, }, end +lemma coe_period : (p : add_circle p) = 0 := +(quotient_add_group.eq_zero_iff p).2 $ mem_zmultiples p + @[simp] lemma coe_add_period (x : 𝕜) : (((x + p) : 𝕜) : add_circle p) = x := -begin - rw [quotient_add_group.coe_add, ←eq_sub_iff_add_eq', sub_self, quotient_add_group.eq_zero_iff], - exact mem_zmultiples p, -end +by rw [coe_add, ←eq_sub_iff_add_eq', sub_self, coe_period] @[continuity, nolint unused_arguments] protected lemma continuous_mk' : continuous (quotient_add_group.mk' (zmultiples p) : 𝕜 → add_circle p) := @@ -195,31 +195,37 @@ section finite_order_points variables {p} -lemma add_order_of_div_of_gcd_eq_one {m n : ℕ} (hn : 0 < n) (h : gcd m n = 1) : - add_order_of (↑(↑m / ↑n * p) : add_circle p) = n := +lemma add_order_of_period_div {n : ℕ} (h : 0 < n) : add_order_of ((p / n : 𝕜) : add_circle p) = n := +begin + rw [add_order_of_eq_iff h], + replace h : 0 < (n : 𝕜) := nat.cast_pos.2 h, + refine ⟨_, λ m hn h0, _⟩; simp only [ne, ← coe_nsmul, nsmul_eq_mul], + { rw [mul_div_cancel' _ h.ne', coe_period] }, + rw coe_eq_zero_of_pos_iff p hp.out (mul_pos (nat.cast_pos.2 h0) $ div_pos hp.out h), + rintro ⟨k, hk⟩, + rw [mul_div, eq_div_iff h.ne', nsmul_eq_mul, mul_right_comm, ← nat.cast_mul, + (mul_left_injective₀ hp.out.ne').eq_iff, nat.cast_inj, mul_comm] at hk, + exact (nat.le_of_dvd h0 ⟨_, hk.symm⟩).not_lt hn, +end + +variables (p) + +lemma gcd_mul_add_order_of_div_eq {n : ℕ} (m : ℕ) (hn : 0 < n) : + m.gcd n * add_order_of (↑(↑m / ↑n * p) : add_circle p) = n := begin - rcases m.eq_zero_or_pos with rfl | hm, { rw [gcd_zero_left, normalize_eq] at h, simp [h], }, - set x : add_circle p := ↑(↑m / ↑n * p), - have hn₀ : (n : 𝕜) ≠ 0, { norm_cast, exact ne_of_gt hn, }, - have hnx : n • x = 0, - { rw [← coe_nsmul, nsmul_eq_mul, ← mul_assoc, mul_div, mul_div_cancel_left _ hn₀, - ← nsmul_eq_mul, quotient_add_group.eq_zero_iff], - exact nsmul_mem_zmultiples p m, }, - apply nat.dvd_antisymm (add_order_of_dvd_of_nsmul_eq_zero hnx), - suffices : ∃ (z : ℕ), z * n = (add_order_of x) * m, - { obtain ⟨z, hz⟩ := this, - simpa only [h, mul_one, gcd_comm n] using dvd_mul_gcd_of_dvd_mul (dvd.intro_left z hz), }, - replace hp := hp.out, - have : 0 < add_order_of x • (↑m / ↑n * p) := smul_pos - (add_order_of_pos' $ (is_of_fin_add_order_iff_nsmul_eq_zero _).2 ⟨n, hn, hnx⟩) (by positivity), - obtain ⟨z, hz⟩ := (coe_eq_zero_of_pos_iff p hp this).mp (add_order_of_nsmul_eq_zero x), - rw [← smul_mul_assoc, nsmul_eq_mul, nsmul_eq_mul, mul_left_inj' hp.ne.symm, mul_div, - eq_div_iff hn₀] at hz, - norm_cast at hz, - exact ⟨z, hz⟩, + rw [mul_comm_div, ← nsmul_eq_mul, coe_nsmul, add_order_of_nsmul''], + { rw [add_order_of_period_div hn, nat.gcd_comm, nat.mul_div_cancel'], + exacts [n.gcd_dvd_left m, hp] }, + { rw [← add_order_of_pos_iff, add_order_of_period_div hn], exacts [hn, hp] }, end -lemma add_order_of_div_of_gcd_eq_one' {m : ℤ} {n : ℕ} (hn : 0 < n) (h : gcd m.nat_abs n = 1) : +variable {p} + +lemma add_order_of_div_of_gcd_eq_one {m n : ℕ} (hn : 0 < n) (h : m.gcd n = 1) : + add_order_of (↑(↑m / ↑n * p) : add_circle p) = n := +by { convert gcd_mul_add_order_of_div_eq p m hn, rw [h, one_mul] } + +lemma add_order_of_div_of_gcd_eq_one' {m : ℤ} {n : ℕ} (hn : 0 < n) (h : m.nat_abs.gcd n = 1) : add_order_of (↑(↑m / ↑n * p) : add_circle p) = n := begin induction m, @@ -237,65 +243,33 @@ begin apply_instance, end -variables (p) - -lemma gcd_mul_add_order_of_div_eq {n : ℕ} (m : ℕ) (hn : 0 < n) : - gcd m n * add_order_of (↑(↑m / ↑n * p) : add_circle p) = n := +lemma add_order_of_eq_pos_iff {u : add_circle p} {n : ℕ} (h : 0 < n) : + add_order_of u = n ↔ ∃ m < n, m.gcd n = 1 ∧ ↑(↑m / ↑n * p) = u := begin - let n' := n / gcd m n, - let m' := m / gcd m n, - have h₀ : 0 < gcd m n, - { rw zero_lt_iff at hn ⊢, contrapose! hn, exact ((gcd_eq_zero_iff m n).mp hn).2, }, - have hk' : 0 < n' := nat.div_pos (nat.le_of_dvd hn $ gcd_dvd_right m n) h₀, - have hgcd : gcd m' n' = 1 := nat.coprime_div_gcd_div_gcd h₀, - simp only [mul_left_inj' hp.out.ne.symm, - ← nat.cast_div_div_div_cancel_right (gcd_dvd_right m n) (gcd_dvd_left m n), - add_order_of_div_of_gcd_eq_one hk' hgcd, mul_comm _ n', nat.div_mul_cancel (gcd_dvd_right m n)], + refine ⟨quotient_add_group.induction_on' u (λ k hk, _), _⟩, swap, + { rintros ⟨m, h₀, h₁, rfl⟩, exact add_order_of_div_of_gcd_eq_one h h₁ }, + have h0 := add_order_of_nsmul_eq_zero (k : add_circle p), + rw [hk, ← coe_nsmul, coe_eq_zero_iff] at h0, + obtain ⟨a, ha⟩ := h0, + have h0 : (_ : 𝕜) ≠ 0 := nat.cast_ne_zero.2 h.ne', + rw [nsmul_eq_mul, mul_comm, ← div_eq_iff h0, ← a.div_add_mod' n, add_smul, add_div, zsmul_eq_mul, + int.cast_mul, int.cast_coe_nat, mul_assoc, ← mul_div, mul_comm _ p, mul_div_cancel p h0] at ha, + have han : _ = a % n := int.to_nat_of_nonneg (int.mod_nonneg _ $ by exact_mod_cast h.ne'), + have he := _, refine ⟨(a % n).to_nat, _, _, he⟩, + { rw [← int.coe_nat_lt, han], + exact int.mod_lt_of_pos _ (int.coe_nat_lt.2 h) }, + { have := (gcd_mul_add_order_of_div_eq p _ h).trans ((congr_arg add_order_of he).trans hk).symm, + rw [he, nat.mul_left_eq_self_iff] at this, { exact this }, { rwa hk } }, + convert congr_arg coe ha using 1, + rw [coe_add, ← int.cast_coe_nat, han, zsmul_eq_mul, mul_div_right_comm, + eq_comm, add_left_eq_self, ← zsmul_eq_mul, coe_zsmul, coe_period, smul_zero], end -variables {p} [floor_ring 𝕜] - lemma exists_gcd_eq_one_of_is_of_fin_add_order {u : add_circle p} (h : is_of_fin_add_order u) : - ∃ m, gcd m (add_order_of u) = 1 ∧ - m < (add_order_of u) ∧ - ↑(((m : 𝕜) / add_order_of u) * p) = u := -begin - rcases eq_or_ne u 0 with rfl | hu, { exact ⟨0, by simp⟩, }, - set n := add_order_of u, - change ∃ m, gcd m n = 1 ∧ m < n ∧ ↑((↑m / ↑n) * p) = u, - have hn : 0 < n := add_order_of_pos' h, - have hn₀ : (n : 𝕜) ≠ 0, { norm_cast, exact ne_of_gt hn, }, - let x := (equiv_Ico p 0 u : 𝕜), - have hxu : (x : add_circle p) = u := (equiv_Ico p 0).symm_apply_apply u, - have hx₀ : 0 < (add_order_of (x : add_circle p)), { rw ← hxu at h, exact add_order_of_pos' h, }, - have hx₁ : 0 < x, - { refine lt_of_le_of_ne (equiv_Ico p 0 u).2.1 _, - contrapose! hu, - rw [← hxu, ← hu, quotient_add_group.coe_zero], }, - obtain ⟨m, hm : m • p = add_order_of ↑x • x⟩ := (coe_eq_zero_of_pos_iff p hp.out - (by positivity)).mp (add_order_of_nsmul_eq_zero (x : add_circle p)), - replace hm : ↑m * p = ↑n * x, { simpa only [hxu, nsmul_eq_mul] using hm, }, - have hux : ↑(↑m / ↑n * p) = u, - { rw [← hxu, ← mul_div_right_comm, hm, mul_comm _ x, mul_div_cancel x hn₀], }, - refine ⟨m, (_ : gcd m n = 1), (_ : m < n), hux⟩, - { have := gcd_mul_add_order_of_div_eq p m hn, - rwa [hux, nat.mul_left_eq_self_iff hn] at this, }, - { have : n • x < n • p := smul_lt_smul_of_pos _ hn, - rwa [nsmul_eq_mul, nsmul_eq_mul, ← hm, mul_lt_mul_right hp.out, nat.cast_lt] at this, - simpa [zero_add] using (equiv_Ico p 0 u).2.2, }, -end - -lemma add_order_of_eq_pos_iff {u : add_circle p} {n : ℕ} (h : 0 < n) : - add_order_of u = n ↔ ∃ m < n, gcd m n = 1 ∧ ↑(↑m / ↑n * p) = u := -begin - refine ⟨λ hu, _, _⟩, - { rw ← hu at h, - obtain ⟨m, h₀, h₁, h₂⟩ := exists_gcd_eq_one_of_is_of_fin_add_order (add_order_of_pos_iff.mp h), - refine ⟨m, _, _, _⟩; - rwa ← hu, }, - { rintros ⟨m, h₀, h₁, rfl⟩, - exact add_order_of_div_of_gcd_eq_one h h₁, }, -end + ∃ m : ℕ, m.gcd (add_order_of u) = 1 ∧ + m < (add_order_of u) ∧ + ↑(((m : 𝕜) / add_order_of u) * p) = u := +let ⟨m, hl, hg, he⟩ := (add_order_of_eq_pos_iff $ add_order_of_pos' h).1 rfl in ⟨m, hg, hl, he⟩ variables (p) @@ -303,31 +277,25 @@ variables (p) `n`. The inverse of the map sends `m ↦ (m/n * p : add_circle p)` where `m` is coprime to `n` and satisfies `0 ≤ m < n`. -/ def set_add_order_of_equiv {n : ℕ} (hn : 0 < n) : - {u : add_circle p | add_order_of u = n} ≃ {m | m < n ∧ gcd m n = 1} := -{ to_fun := λ u, by - { let h := (add_order_of_eq_pos_iff hn).mp u.property, - exact ⟨classical.some h, classical.some (classical.some_spec h), - (classical.some_spec (classical.some_spec h)).1⟩, }, - inv_fun := λ m, ⟨↑((m : 𝕜) / n * p), add_order_of_div_of_gcd_eq_one hn (m.property.2)⟩, - left_inv := λ u, subtype.ext - (classical.some_spec (classical.some_spec $ (add_order_of_eq_pos_iff hn).mp u.2)).2, - right_inv := - begin - rintros ⟨m, hm₁, hm₂⟩, - let u : {u : add_circle p | add_order_of u = n} := - ⟨↑((m : 𝕜) / n * p), add_order_of_div_of_gcd_eq_one hn hm₂⟩, - let h := (add_order_of_eq_pos_iff hn).mp u.property, - ext, - let m' := classical.some h, - change m' = m, - obtain ⟨h₁ : m' < n, h₂ : gcd m' n = 1, h₃ : quotient_add_group.mk ((m' : 𝕜) / n * p) = - quotient_add_group.mk ((m : 𝕜) / n * p)⟩ := classical.some_spec h, - replace h₃ := congr_arg (coe : Ico 0 (0 + p) → 𝕜) (congr_arg (equiv_Ico p 0) h₃), - simpa only [coe_equiv_Ico_mk_apply, mul_left_inj' hp.out.ne', mul_div_cancel _ hp.out.ne', - int.fract_div_nat_cast_eq_div_nat_cast_mod, - div_left_inj' (nat.cast_ne_zero.mpr hn.ne' : (n : 𝕜) ≠ 0), nat.cast_inj, - (nat.mod_eq_iff_lt hn.ne').mpr hm₁, (nat.mod_eq_iff_lt hn.ne').mpr h₁] using h₃, - end } + {u : add_circle p | add_order_of u = n} ≃ {m | m < n ∧ m.gcd n = 1} := +equiv.symm $ equiv.of_bijective + (λ m, ⟨↑((m : 𝕜) / n * p), add_order_of_div_of_gcd_eq_one hn (m.prop.2)⟩) +begin + refine ⟨λ m₁ m₂ h, subtype.ext _, λ u, _⟩, + { simp_rw [subtype.ext_iff, subtype.coe_mk] at h, + rw [← sub_eq_zero, ← coe_sub, ← sub_mul, ← sub_div, coe_coe, coe_coe, ← int.cast_coe_nat m₁, + ← int.cast_coe_nat m₂, ← int.cast_sub, coe_eq_zero_iff] at h, + obtain ⟨m, hm⟩ := h, + rw [← mul_div_right_comm, eq_div_iff, mul_comm, ← zsmul_eq_mul, mul_smul_comm, ← nsmul_eq_mul, + ← coe_nat_zsmul, smul_smul, (zsmul_strict_mono_left hp.out).injective.eq_iff, mul_comm] at hm, + swap, { exact nat.cast_ne_zero.2 hn.ne' }, + rw [← @nat.cast_inj ℤ, ← sub_eq_zero], + refine int.eq_zero_of_abs_lt_dvd ⟨_, hm.symm⟩ (abs_sub_lt_iff.2 ⟨_, _⟩); + apply (int.sub_le_self _ $ nat.cast_nonneg _).trans_lt (nat.cast_lt.2 _), + exacts [m₁.2.1, m₂.2.1] }, + obtain ⟨m, hmn, hg, he⟩ := (add_order_of_eq_pos_iff hn).mp u.2, + exact ⟨⟨m, hmn, hg⟩, subtype.ext he⟩, +end @[simp] lemma card_add_order_of_eq_totient {n : ℕ} : nat.card {u : add_circle p // add_order_of u = n} = n.totient := @@ -343,7 +311,7 @@ begin exact nat.card_of_is_empty, }, }, { rw [← coe_set_of, nat.card_congr (set_add_order_of_equiv p hn), n.totient_eq_card_lt_and_coprime], - simpa only [@nat.coprime_comm _ n], }, + simp only [nat.gcd_comm], }, end lemma finite_set_of_add_order_eq {n : ℕ} (hn : 0 < n) : From f61ffad4cfd25f859269f01d7383b73e8f6feb90 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 31 Dec 2022 05:31:36 +0000 Subject: [PATCH 0118/1291] chore(group_theory/coset): Additivise `quotient_group.preimage_mk_equiv_subgroup_times_set` (#18008) The matches in the data fields were tripping `to_additive` up. Make one more argument to `mk_mul_of_mem`, fix the additive lemma name and remove a now useless argument to `card_quotient_dvd_card`. --- src/group_theory/coset.lean | 46 +++++++++++++++---------------------- 1 file changed, 19 insertions(+), 27 deletions(-) diff --git a/src/group_theory/coset.lean b/src/group_theory/coset.lean index e89d65f50d697..48d24dd1d9afa 100644 --- a/src/group_theory/coset.lean +++ b/src/group_theory/coset.lean @@ -34,10 +34,6 @@ This file develops the basic theory of left and right cosets. * `s +r a`: for `right_add_coset s a`. * `G ⧸ H` is the quotient of the (additive) group `G` by the (additive) subgroup `H` - -## TODO - -Add `to_additive` to `preimage_mk_equiv_subgroup_times_set`. -/ open set function @@ -363,7 +359,7 @@ calc _ ↔ @setoid.r _ (left_rel s) a b : quotient.eq' lemma eq' {a b : α} : (mk a : α ⧸ s) = mk b ↔ a⁻¹ * b ∈ s := quotient_group.eq -@[to_additive quotient_add_group.out_eq'] +@[simp, to_additive quotient_add_group.out_eq'] lemma out_eq' (a : α ⧸ s) : mk a.out' = a := quotient.out_eq' a @@ -376,10 +372,10 @@ variables (s) lemma mk_out'_eq_mul (g : α) : ∃ h : s, (mk g : α ⧸ s).out' = g * h := ⟨⟨g⁻¹ * (mk g).out', eq'.mp (mk g).out_eq'.symm⟩, by rw [set_like.coe_mk, mul_inv_cancel_left]⟩ -variables {s} +variables {s} {a b : α} -@[to_additive quotient_add_group.mk_mul_of_mem] -lemma mk_mul_of_mem (g₁ g₂ : α) (hg₂ : g₂ ∈ s) : (mk (g₁ * g₂) : α ⧸ s) = mk g₁ := +@[simp, to_additive quotient_add_group.mk_add_of_mem] +lemma mk_mul_of_mem (a : α) (hb : b ∈ s) : (mk (a * b) : α ⧸ s) = mk a := by rwa [eq', mul_inv_rev, inv_mul_cancel_right, s.inv_mem_iff] @[to_additive] @@ -481,7 +477,7 @@ def quotient_equiv_prod_of_le' (h_le : s ≤ t) refine quotient.ind' (λ a, _), refine quotient.ind' (λ b, _), have key : quotient.mk' (f (quotient.mk' a) * b) = quotient.mk' a := - (quotient_group.mk_mul_of_mem (f a) ↑b b.2).trans (hf a), + (quotient_group.mk_mul_of_mem (f a) b.2).trans (hf a), simp_rw [quotient.map'_mk', id.def, key, inv_mul_cancel_left, subtype.coe_eta] } } /-- If `H ≤ K`, then `G/H ≃ G/K × K/H` nonconstructively. @@ -567,8 +563,8 @@ lemma card_subgroup_dvd_card [fintype α] (s : subgroup α) [fintype s] : fintype.card s ∣ fintype.card α := by classical; simp [card_eq_card_quotient_mul_card_subgroup s, @dvd_mul_left ℕ] -@[to_additive] lemma card_quotient_dvd_card [fintype α] (s : subgroup α) - [decidable_pred (λ a, a ∈ s)] [fintype s] : fintype.card (α ⧸ s) ∣ fintype.card α := +@[to_additive] lemma card_quotient_dvd_card [fintype α] (s : subgroup α) [decidable_pred (∈ s)] : + fintype.card (α ⧸ s) ∣ fintype.card α := by simp [card_eq_card_quotient_mul_card_subgroup s, @dvd_mul_right ℕ] open fintype @@ -600,24 +596,20 @@ namespace quotient_group variables [group α] --- FIXME -- why is there no `to_additive`? - -/-- If `s` is a subgroup of the group `α`, and `t` is a subset of `α/s`, then -there is a (typically non-canonical) bijection between the preimage of `t` in -`α` and the product `s × t`. -/ -noncomputable def preimage_mk_equiv_subgroup_times_set - (s : subgroup α) (t : set (α ⧸ s)) : quotient_group.mk ⁻¹' t ≃ s × t := -have h : ∀ {x : α ⧸ s} {a : α}, x ∈ t → a ∈ s → - (quotient.mk' (quotient.out' x * a) : α ⧸ s) = quotient.mk' (quotient.out' x) := - λ x a hx ha, quotient.sound' $ by rwa [left_rel_apply, ← s.inv_mem_iff, mul_inv_rev, inv_inv, - ← mul_assoc, inv_mul_self, one_mul], -{ to_fun := λ ⟨a, ha⟩, ⟨⟨(quotient.out' (quotient.mk' a))⁻¹ * a, +/-- If `s` is a subgroup of the group `α`, and `t` is a subset of `α ⧸ s`, then there is a +(typically non-canonical) bijection between the preimage of `t` in `α` and the product `s × t`. -/ +@[to_additive "If `s` is a subgroup of the additive group `α`, and `t` is a subset of `α ⧸ s`, then +there is a (typically non-canonical) bijection between the preimage of `t` in `α` and the product +`s × t`."] +noncomputable def preimage_mk_equiv_subgroup_times_set (s : subgroup α) (t : set (α ⧸ s)) : + quotient_group.mk ⁻¹' t ≃ s × t := +{ to_fun := λ a, ⟨⟨(quotient.out' (quotient_group.mk a))⁻¹ * a, left_rel_apply.mp (@quotient.exact' _ (left_rel s) _ _ $ (quotient.out_eq' _))⟩, - ⟨quotient.mk' a, ha⟩⟩, - inv_fun := λ ⟨⟨a, ha⟩, ⟨x, hx⟩⟩, ⟨quotient.out' x * a, show quotient.mk' _ ∈ t, - by simp [h hx ha, hx]⟩, + ⟨quotient_group.mk a, a.2⟩⟩, + inv_fun := λ a, ⟨quotient.out' a.2.1 * a.1.1, show quotient_group.mk _ ∈ t, + by { rw [mk_mul_of_mem _ a.1.2, out_eq'], exact a.2.2 }⟩, left_inv := λ ⟨a, ha⟩, subtype.eq $ show _ * _ = a, by simp, - right_inv := λ ⟨⟨a, ha⟩, ⟨x, hx⟩⟩, show (_, _) = _, by simp [h hx ha] } + right_inv := λ ⟨⟨a, ha⟩, ⟨x, hx⟩⟩, by ext; simp [ha] } end quotient_group From 14e80e85cbca5872a329fbfd3d1f3fd64e306934 Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Sat, 31 Dec 2022 09:18:46 +0000 Subject: [PATCH 0119/1291] feat(category_theory/filtered): generalize to allow empty categories (#18013) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Also: + Create aliases so that we can write `is_filtered.cocone_objs` without the `_or_empty`. + Fix misnomers `is_cofiltered.cocone_objs/maps`; `cocone` should be replaced by `cone`. + Add special shapes `span`/`cospan` and lemma `ranges_directed` from #17905. + Golf `bowtie` and `tulip` using `span`. Co-authored-by: Rémi Bottinelli --- src/category_theory/filtered.lean | 177 ++++++++++++++------------ src/category_theory/functor/flat.lean | 4 +- 2 files changed, 97 insertions(+), 84 deletions(-) diff --git a/src/category_theory/filtered.lean b/src/category_theory/filtered.lean index c430450c72be7..2dfb410c9718d 100644 --- a/src/category_theory/filtered.lean +++ b/src/category_theory/filtered.lean @@ -97,7 +97,7 @@ instance is_filtered_or_empty_of_directed_le (α : Type u) [preorder α] [is_dir cocone_maps := λ X Y f g, ⟨Y, 𝟙 _, by simp⟩ } @[priority 100] -instance is_filtered_of_directed_le_nonempty (α : Type u) [preorder α] [is_directed α (≤)] +instance is_filtered_of_directed_le_nonempty (α : Type u) [preorder α] [is_directed α (≤)] [nonempty α] : is_filtered α := {} @@ -112,28 +112,35 @@ instance : is_filtered (discrete punit) := namespace is_filtered -variables {C} [is_filtered C] +section allow_empty + +variables {C} [is_filtered_or_empty C] + +lemma cocone_objs : ∀ (X Y : C), ∃ Z (f : X ⟶ Z) (g : Y ⟶ Z), true := +is_filtered_or_empty.cocone_objs +lemma cocone_maps : ∀ ⦃X Y : C⦄ (f g : X ⟶ Y), ∃ Z (h : Y ⟶ Z), f ≫ h = g ≫ h := +is_filtered_or_empty.cocone_maps /-- `max j j'` is an arbitrary choice of object to the right of both `j` and `j'`, whose existence is ensured by `is_filtered`. -/ noncomputable def max (j j' : C) : C := -(is_filtered_or_empty.cocone_objs j j').some +(cocone_objs j j').some /-- -`left_to_max j j'` is an arbitrarily choice of morphism from `j` to `max j j'`, +`left_to_max j j'` is an arbitrary choice of morphism from `j` to `max j j'`, whose existence is ensured by `is_filtered`. -/ noncomputable def left_to_max (j j' : C) : j ⟶ max j j' := -(is_filtered_or_empty.cocone_objs j j').some_spec.some +(cocone_objs j j').some_spec.some /-- -`right_to_max j j'` is an arbitrarily choice of morphism from `j'` to `max j j'`, +`right_to_max j j'` is an arbitrary choice of morphism from `j'` to `max j j'`, whose existence is ensured by `is_filtered`. -/ noncomputable def right_to_max (j j' : C) : j' ⟶ max j j' := -(is_filtered_or_empty.cocone_objs j j').some_spec.some_spec.some +(cocone_objs j j').some_spec.some_spec.some /-- `coeq f f'`, for morphisms `f f' : j ⟶ j'`, is an arbitrary choice of object @@ -142,7 +149,7 @@ which admits a morphism `coeq_hom f f' : j' ⟶ coeq f f'` such that Its existence is ensured by `is_filtered`. -/ noncomputable def coeq {j j' : C} (f f' : j ⟶ j') : C := -(is_filtered_or_empty.cocone_maps f f').some +(cocone_maps f f').some /-- `coeq_hom f f'`, for morphisms `f f' : j ⟶ j'`, is an arbitrary choice of morphism @@ -151,7 +158,7 @@ noncomputable def coeq {j j' : C} (f f' : j ⟶ j') : C := Its existence is ensured by `is_filtered`. -/ noncomputable def coeq_hom {j j' : C} (f f' : j ⟶ j') : j' ⟶ coeq f f' := -(is_filtered_or_empty.cocone_maps f f').some_spec.some +(cocone_maps f f').some_spec.some /-- `coeq_condition f f'`, for morphisms `f f' : j ⟶ j'`, is the proof that @@ -159,10 +166,16 @@ noncomputable def coeq_hom {j j' : C} (f f' : j ⟶ j') : j' ⟶ coeq f f' := -/ @[simp, reassoc] lemma coeq_condition {j j' : C} (f f' : j ⟶ j') : f ≫ coeq_hom f f' = f' ≫ coeq_hom f f' := -(is_filtered_or_empty.cocone_maps f f').some_spec.some_spec +(cocone_maps f f').some_spec.some_spec + +end allow_empty + +section nonempty open category_theory.limits +variables {C} [is_filtered C] + /-- Any finite collection of objects in a filtered category has an object "to the right". -/ @@ -291,8 +304,12 @@ of_right_adjoint (adjunction.of_right_adjoint R) lemma of_equivalence (h : C ≌ D) : is_filtered D := of_right_adjoint h.symm.to_adjunction +end nonempty + section special_shapes +variables {C} [is_filtered_or_empty C] + /-- `max₃ j₁ j₂ j₃` is an arbitrary choice of object to the right of `j₁`, `j₂` and `j₃`, whose existence is ensured by `is_filtered`. @@ -300,21 +317,21 @@ whose existence is ensured by `is_filtered`. noncomputable def max₃ (j₁ j₂ j₃ : C) : C := max (max j₁ j₂) j₃ /-- -`first_to_max₃ j₁ j₂ j₃` is an arbitrarily choice of morphism from `j₁` to `max₃ j₁ j₂ j₃`, +`first_to_max₃ j₁ j₂ j₃` is an arbitrary choice of morphism from `j₁` to `max₃ j₁ j₂ j₃`, whose existence is ensured by `is_filtered`. -/ noncomputable def first_to_max₃ (j₁ j₂ j₃ : C) : j₁ ⟶ max₃ j₁ j₂ j₃ := left_to_max j₁ j₂ ≫ left_to_max (max j₁ j₂) j₃ /-- -`second_to_max₃ j₁ j₂ j₃` is an arbitrarily choice of morphism from `j₂` to `max₃ j₁ j₂ j₃`, +`second_to_max₃ j₁ j₂ j₃` is an arbitrary choice of morphism from `j₂` to `max₃ j₁ j₂ j₃`, whose existence is ensured by `is_filtered`. -/ noncomputable def second_to_max₃ (j₁ j₂ j₃ : C) : j₂ ⟶ max₃ j₁ j₂ j₃ := right_to_max j₁ j₂ ≫ left_to_max (max j₁ j₂) j₃ /-- -`third_to_max₃ j₁ j₂ j₃` is an arbitrarily choice of morphism from `j₃` to `max₃ j₁ j₂ j₃`, +`third_to_max₃ j₁ j₂ j₃` is an arbitrary choice of morphism from `j₃` to `max₃ j₁ j₂ j₃`, whose existence is ensured by `is_filtered`. -/ noncomputable def third_to_max₃ (j₁ j₂ j₃ : C) : j₃ ⟶ max₃ j₁ j₂ j₃ := @@ -342,11 +359,7 @@ coeq_hom (coeq_hom f g ≫ left_to_max (coeq f g) (coeq g h)) lemma coeq₃_condition₁ {j₁ j₂ : C} (f g h : j₁ ⟶ j₂) : f ≫ coeq₃_hom f g h = g ≫ coeq₃_hom f g h := -begin - dsimp [coeq₃_hom], - slice_lhs 1 2 { rw coeq_condition f g }, - simp only [category.assoc], -end +by rw [coeq₃_hom, reassoc_of (coeq_condition f g)] lemma coeq₃_condition₂ {j₁ j₂ : C} (f g h : j₁ ⟶ j₂) : g ≫ coeq₃_hom f g h = h ≫ coeq₃_hom f g h := @@ -362,6 +375,13 @@ lemma coeq₃_condition₃ {j₁ j₂ : C} (f g h : j₁ ⟶ j₂) : f ≫ coeq₃_hom f g h = h ≫ coeq₃_hom f g h := eq.trans (coeq₃_condition₁ f g h) (coeq₃_condition₂ f g h) +/-- For every span `j ⟵ i ⟶ j'`, there + exists a cocone `j ⟶ k ⟵ j'` such that the square commutes. -/ +lemma span {i j j' : C} (f : i ⟶ j) (f' : i ⟶ j') : + ∃ (k : C) (g : j ⟶ k) (g' : j' ⟶ k), f ≫ g = f' ≫ g' := +let ⟨K, G, G', _⟩ := cocone_objs j j', ⟨k, e, he⟩ := cocone_maps (f ≫ G) (f' ≫ G') in +⟨k, G ≫ e, G' ≫ e, by simpa only [← category.assoc]⟩ + /-- Given a "bowtie" of morphisms ``` @@ -380,24 +400,10 @@ lemma bowtie {j₁ j₂ k₁ k₂ : C} (f₁ : j₁ ⟶ k₁) (g₁ : j₁ ⟶ k₂) (f₂ : j₂ ⟶ k₁) (g₂ : j₂ ⟶ k₂) : ∃ (s : C) (α : k₁ ⟶ s) (β : k₂ ⟶ s), f₁ ≫ α = g₁ ≫ β ∧ f₂ ≫ α = g₂ ≫ β := begin - let sa := max k₁ k₂, - let sb := coeq (f₁ ≫ left_to_max _ _) (g₁ ≫ right_to_max _ _), - let sc := coeq (f₂ ≫ left_to_max _ _) (g₂ ≫ right_to_max _ _), - let sd := max sb sc, - let s := coeq ((coeq_hom _ _ : sa ⟶ sb) ≫ left_to_max _ _) - ((coeq_hom _ _ : sa ⟶ sc) ≫ right_to_max _ _), - use s, - fsplit, - exact left_to_max k₁ k₂ ≫ coeq_hom _ _ ≫ left_to_max sb sc ≫ coeq_hom _ _, - fsplit, - exact right_to_max k₁ k₂ ≫ coeq_hom _ _ ≫ right_to_max sb sc ≫ coeq_hom _ _, - fsplit, - { slice_lhs 1 3 { rw [←category.assoc, coeq_condition], }, - slice_lhs 3 5 { rw [←category.assoc, coeq_condition], }, - simp only [category.assoc], }, - { slice_lhs 3 5 { rw [←category.assoc, coeq_condition], }, - slice_lhs 1 3 { rw [←category.assoc, coeq_condition], }, - simp only [category.assoc], } + obtain ⟨t, k₁t, k₂t, ht⟩ := span f₁ g₁, + obtain ⟨s, ts, hs⟩ := cocone_maps (f₂ ≫ k₁t) (g₂ ≫ k₂t), + simp_rw category.assoc at hs, + exact ⟨s, k₁t ≫ ts, k₂t ≫ ts, by rw reassoc_of ht, hs⟩, end /-- @@ -416,36 +422,17 @@ Given a "tulip" of morphisms l ``` in a filtered category, we can construct an object `s` and three morphisms from `k₁`, `k₂` and `l` -to `s`, making the resulting sqaures commute. +to `s`, making the resulting squares commute. -/ lemma tulip {j₁ j₂ j₃ k₁ k₂ l : C} (f₁ : j₁ ⟶ k₁) (f₂ : j₂ ⟶ k₁) (f₃ : j₂ ⟶ k₂) (f₄ : j₃ ⟶ k₂) (g₁ : j₁ ⟶ l) (g₂ : j₃ ⟶ l) : ∃ (s : C) (α : k₁ ⟶ s) (β : l ⟶ s) (γ : k₂ ⟶ s), f₁ ≫ α = g₁ ≫ β ∧ f₂ ≫ α = f₃ ≫ γ ∧ f₄ ≫ γ = g₂ ≫ β := begin - let sa := max₃ k₁ l k₂, - let sb := coeq (f₁ ≫ first_to_max₃ k₁ l k₂) (g₁ ≫ second_to_max₃ k₁ l k₂), - let sc := coeq (f₂ ≫ first_to_max₃ k₁ l k₂) (f₃ ≫ third_to_max₃ k₁ l k₂), - let sd := coeq (f₄ ≫ third_to_max₃ k₁ l k₂) (g₂ ≫ second_to_max₃ k₁ l k₂), - let se := max₃ sb sc sd, - let sf := coeq₃ (coeq_hom _ _ ≫ first_to_max₃ sb sc sd) - (coeq_hom _ _ ≫ second_to_max₃ sb sc sd) (coeq_hom _ _ ≫ third_to_max₃ sb sc sd), - use sf, - use first_to_max₃ k₁ l k₂ ≫ coeq_hom _ _ ≫ first_to_max₃ sb sc sd ≫ coeq₃_hom _ _ _, - use second_to_max₃ k₁ l k₂ ≫ coeq_hom _ _ ≫ second_to_max₃ sb sc sd ≫ coeq₃_hom _ _ _, - use third_to_max₃ k₁ l k₂ ≫ coeq_hom _ _ ≫ third_to_max₃ sb sc sd ≫ coeq₃_hom _ _ _, - fsplit, - slice_lhs 1 3 { rw [← category.assoc, coeq_condition] }, - slice_lhs 3 6 { rw [← category.assoc, coeq₃_condition₁] }, - simp only [category.assoc], - fsplit, - slice_lhs 3 6 { rw [← category.assoc, coeq₃_condition₁] }, - slice_lhs 1 3 { rw [← category.assoc, coeq_condition] }, - slice_rhs 3 6 { rw [← category.assoc, ← coeq₃_condition₂] }, - simp only [category.assoc], - slice_rhs 3 6 { rw [← category.assoc, coeq₃_condition₂] }, - slice_rhs 1 3 { rw [← category.assoc, ← coeq_condition] }, - simp only [category.assoc], + obtain ⟨l', k₁l, k₂l, hl⟩ := span f₂ f₃, + obtain ⟨s, ls, l's, hs₁, hs₂⟩ := bowtie g₁ (f₁ ≫ k₁l) g₂ (f₄ ≫ k₂l), + refine ⟨s, k₁l ≫ l's, ls, k₂l ≫ l's, _, by rw reassoc_of hl, _⟩; + simp only [hs₁, hs₂, category.assoc], end end special_shapes @@ -459,8 +446,8 @@ A category `is_cofiltered_or_empty` if are equal. -/ class is_cofiltered_or_empty : Prop := -(cocone_objs : ∀ (X Y : C), ∃ W (f : W ⟶ X) (g : W ⟶ Y), true) -(cocone_maps : ∀ ⦃X Y : C⦄ (f g : X ⟶ Y), ∃ W (h : W ⟶ X), h ≫ f = h ≫ g) +(cone_objs : ∀ (X Y : C), ∃ W (f : W ⟶ X) (g : W ⟶ Y), true) +(cone_maps : ∀ ⦃X Y : C⦄ (f g : X ⟶ Y), ∃ W (h : W ⟶ X), h ≫ f = h ≫ g) /-- A category `is_cofiltered` if @@ -477,8 +464,8 @@ class is_cofiltered extends is_cofiltered_or_empty C : Prop := @[priority 100] instance is_cofiltered_or_empty_of_semilattice_inf (α : Type u) [semilattice_inf α] : is_cofiltered_or_empty α := -{ cocone_objs := λ X Y, ⟨X ⊓ Y, hom_of_le inf_le_left, hom_of_le inf_le_right, trivial⟩, - cocone_maps := λ X Y f g, ⟨X, 𝟙 _, (by ext)⟩, } +{ cone_objs := λ X Y, ⟨X ⊓ Y, hom_of_le inf_le_left, hom_of_le inf_le_right, trivial⟩, + cone_maps := λ X Y f g, ⟨X, 𝟙 _, (by ext)⟩, } @[priority 100] instance is_cofiltered_of_semilattice_inf_nonempty @@ -488,12 +475,12 @@ instance is_cofiltered_of_semilattice_inf_nonempty instance is_cofiltered_or_empty_of_directed_ge (α : Type u) [preorder α] [is_directed α (≥)] : is_cofiltered_or_empty α := -{ cocone_objs := λ X Y, let ⟨Z, hX, hY⟩ := exists_le_le X Y in +{ cone_objs := λ X Y, let ⟨Z, hX, hY⟩ := exists_le_le X Y in ⟨Z, hom_of_le hX, hom_of_le hY, trivial⟩, - cocone_maps := λ X Y f g, ⟨X, 𝟙 _, by simp⟩ } + cone_maps := λ X Y f g, ⟨X, 𝟙 _, by simp⟩ } @[priority 100] -instance is_cofiltered_of_directed_ge_nonempty (α : Type u) [preorder α] [is_directed α (≥)] +instance is_cofiltered_of_directed_ge_nonempty (α : Type u) [preorder α] [is_directed α (≥)] [nonempty α] : is_cofiltered α := {} @@ -502,34 +489,40 @@ example (α : Type u) [semilattice_inf α] [order_bot α] : is_cofiltered α := example (α : Type u) [semilattice_inf α] [order_top α] : is_cofiltered α := by apply_instance instance : is_cofiltered (discrete punit) := -{ cocone_objs := λ X Y, ⟨⟨punit.star⟩, ⟨⟨dec_trivial⟩⟩, ⟨⟨dec_trivial⟩⟩, trivial⟩, - cocone_maps := λ X Y f g, ⟨⟨punit.star⟩, ⟨⟨dec_trivial⟩⟩, dec_trivial⟩, +{ cone_objs := λ X Y, ⟨⟨punit.star⟩, ⟨⟨dec_trivial⟩⟩, ⟨⟨dec_trivial⟩⟩, trivial⟩, + cone_maps := λ X Y f g, ⟨⟨punit.star⟩, ⟨⟨dec_trivial⟩⟩, dec_trivial⟩, nonempty := ⟨⟨punit.star⟩⟩ } namespace is_cofiltered -variables {C} [is_cofiltered C] +section allow_empty + +variables {C} [is_cofiltered_or_empty C] + +lemma cone_objs : ∀ (X Y : C), ∃ W (f : W ⟶ X) (g : W ⟶ Y), true := is_cofiltered_or_empty.cone_objs +lemma cone_maps : ∀ ⦃X Y : C⦄ (f g : X ⟶ Y), ∃ W (h : W ⟶ X), h ≫ f = h ≫ g := +is_cofiltered_or_empty.cone_maps /-- `min j j'` is an arbitrary choice of object to the left of both `j` and `j'`, whose existence is ensured by `is_cofiltered`. -/ noncomputable def min (j j' : C) : C := -(is_cofiltered_or_empty.cocone_objs j j').some +(cone_objs j j').some /-- -`min_to_left j j'` is an arbitrarily choice of morphism from `min j j'` to `j`, +`min_to_left j j'` is an arbitrary choice of morphism from `min j j'` to `j`, whose existence is ensured by `is_cofiltered`. -/ noncomputable def min_to_left (j j' : C) : min j j' ⟶ j := -(is_cofiltered_or_empty.cocone_objs j j').some_spec.some +(cone_objs j j').some_spec.some /-- -`min_to_right j j'` is an arbitrarily choice of morphism from `min j j'` to `j'`, +`min_to_right j j'` is an arbitrary choice of morphism from `min j j'` to `j'`, whose existence is ensured by `is_cofiltered`. -/ noncomputable def min_to_right (j j' : C) : min j j' ⟶ j' := -(is_cofiltered_or_empty.cocone_objs j j').some_spec.some_spec.some +(cone_objs j j').some_spec.some_spec.some /-- `eq f f'`, for morphisms `f f' : j ⟶ j'`, is an arbitrary choice of object @@ -538,7 +531,7 @@ which admits a morphism `eq_hom f f' : eq f f' ⟶ j` such that Its existence is ensured by `is_cofiltered`. -/ noncomputable def eq {j j' : C} (f f' : j ⟶ j') : C := -(is_cofiltered_or_empty.cocone_maps f f').some +(cone_maps f f').some /-- `eq_hom f f'`, for morphisms `f f' : j ⟶ j'`, is an arbitrary choice of morphism @@ -547,7 +540,7 @@ noncomputable def eq {j j' : C} (f f' : j ⟶ j') : C := Its existence is ensured by `is_cofiltered`. -/ noncomputable def eq_hom {j j' : C} (f f' : j ⟶ j') : eq f f' ⟶ j := -(is_cofiltered_or_empty.cocone_maps f f').some_spec.some +(cone_maps f f').some_spec.some /-- `eq_condition f f'`, for morphisms `f f' : j ⟶ j'`, is the proof that @@ -555,10 +548,28 @@ noncomputable def eq_hom {j j' : C} (f f' : j ⟶ j') : eq f f' ⟶ j := -/ @[simp, reassoc] lemma eq_condition {j j' : C} (f f' : j ⟶ j') : eq_hom f f' ≫ f = eq_hom f f' ≫ f' := -(is_cofiltered_or_empty.cocone_maps f f').some_spec.some_spec +(cone_maps f f').some_spec.some_spec + +/-- For every cospan `j ⟶ i ⟵ j'`, + there exists a cone `j ⟵ k ⟶ j'` such that the square commutes. -/ +lemma cospan {i j j' : C} (f : j ⟶ i) (f' : j' ⟶ i) : + ∃ (k : C) (g : k ⟶ j) (g' : k ⟶ j'), g ≫ f = g' ≫ f' := +let ⟨K, G, G', _⟩ := cone_objs j j', ⟨k, e, he⟩ := cone_maps (G ≫ f) (G' ≫ f') in +⟨k, e ≫ G, e ≫ G', by simpa only [category.assoc] using he⟩ + +lemma _root_.category_theory.functor.ranges_directed (F : C ⥤ Type*) (j : C) : + directed (⊇) (λ (f : Σ' i, i ⟶ j), set.range (F.map f.2)) := +λ ⟨i, ij⟩ ⟨k, kj⟩, let ⟨l, li, lk, e⟩ := cospan ij kj in +by refine ⟨⟨l, lk ≫ kj⟩, e ▸ _, _⟩; simp_rw F.map_comp; apply set.range_comp_subset_range + +end allow_empty + +section nonempty open category_theory.limits +variables {C} [is_cofiltered C] + /-- Any finite collection of objects in a cofiltered category has an object "to the left". -/ @@ -674,10 +685,10 @@ If `C` is cofiltered, and we have a functor `L : C ⥤ D` with a right adjoint, then `D` is cofiltered. -/ lemma of_left_adjoint {L : C ⥤ D} {R : D ⥤ C} (h : L ⊣ R) : is_cofiltered D := -{ cocone_objs := λ X Y, +{ cone_objs := λ X Y, ⟨L.obj (min (R.obj X) (R.obj Y)), (h.hom_equiv _ X).symm (min_to_left _ _), (h.hom_equiv _ Y).symm (min_to_right _ _), ⟨⟩⟩, - cocone_maps := λ X Y f g, + cone_maps := λ X Y f g, ⟨L.obj (eq (R.map f) (R.map g)), (h.hom_equiv _ _).symm (eq_hom _ _), by rw [← h.hom_equiv_naturality_right_symm, ← h.hom_equiv_naturality_right_symm, eq_condition]⟩, @@ -691,15 +702,17 @@ of_left_adjoint (adjunction.of_left_adjoint L) lemma of_equivalence (h : C ≌ D) : is_cofiltered D := of_left_adjoint h.to_adjunction +end nonempty + end is_cofiltered section opposite open opposite instance is_cofiltered_op_of_is_filtered [is_filtered C] : is_cofiltered Cᵒᵖ := -{ cocone_objs := λ X Y, ⟨op (is_filtered.max X.unop Y.unop), +{ cone_objs := λ X Y, ⟨op (is_filtered.max X.unop Y.unop), (is_filtered.left_to_max _ _).op, (is_filtered.right_to_max _ _).op, trivial⟩, - cocone_maps := λ X Y f g, ⟨op (is_filtered.coeq f.unop g.unop), + cone_maps := λ X Y f g, ⟨op (is_filtered.coeq f.unop g.unop), (is_filtered.coeq_hom _ _).op, begin rw [(show f = f.unop.op, by simp), (show g = g.unop.op, by simp), ← op_comp, ← op_comp], diff --git a/src/category_theory/functor/flat.lean b/src/category_theory/functor/flat.lean index 90f129c513cfb..b8adef71d1b89 100644 --- a/src/category_theory/functor/flat.lean +++ b/src/category_theory/functor/flat.lean @@ -169,8 +169,8 @@ variables {C : Type u₁} [category.{v₁} C] {D : Type u₂} [category.{v₁} D local attribute [instance] has_finite_limits_of_has_finite_limits_of_size lemma cofiltered_of_has_finite_limits [has_finite_limits C] : is_cofiltered C := -{ cocone_objs := λ A B, ⟨limits.prod A B, limits.prod.fst, limits.prod.snd, trivial⟩, - cocone_maps := λ A B f g, ⟨equalizer f g, equalizer.ι f g, equalizer.condition f g⟩, +{ cone_objs := λ A B, ⟨limits.prod A B, limits.prod.fst, limits.prod.snd, trivial⟩, + cone_maps := λ A B f g, ⟨equalizer f g, equalizer.ι f g, equalizer.condition f g⟩, nonempty := ⟨⊤_ C⟩ } lemma flat_of_preserves_finite_limits [has_finite_limits C] (F : C ⥤ D) From ffc3730d545623aedf5d5bd46a3153cbf41f6c2c Mon Sep 17 00:00:00 2001 From: Oliver Nash Date: Sat, 31 Dec 2022 14:21:55 +0000 Subject: [PATCH 0120/1291] feat(number_theory/well_approximable): well-approximable numbers and Gallagher's zero-one law (#18011) --- docs/references.bib | 33 ++ src/analysis/normed/group/basic.lean | 10 + .../covering/liminf_limsup.lean | 66 ++++ src/number_theory/well_approximable.lean | 321 ++++++++++++++++++ src/order/liminf_limsup.lean | 17 + .../metric_space/hausdorff_distance.lean | 3 + 6 files changed, 450 insertions(+) create mode 100644 src/number_theory/well_approximable.lean diff --git a/docs/references.bib b/docs/references.bib index 8620c56acc625..47a3cec0899f5 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -899,6 +899,22 @@ @Book{ gabriel-zisman-1967 pages = {x+168} } +@Article{ Gallagher1961, + author = {Gallagher, Patrick}, + title = {Approximation by reduced fractions}, + journal = {J. Math. Soc. Japan}, + fjournal = {Journal of the Mathematical Society of Japan}, + volume = {13}, + year = {1961}, + pages = {342--345}, + issn = {0025-5645}, + mrclass = {10.30}, + mrnumber = {133297}, + mrreviewer = {J. W. S. Cassels}, + doi = {10.2969/jmsj/01340342}, + url = {https://doi.org/10.2969/jmsj/01340342}, +} + @InProceedings{ Gallier2011Notes, title = {Notes on Differential Geometry and Lie Groups}, author = {J. Gallier and J. Quaintance}, @@ -1363,6 +1379,23 @@ @Article{ kleitman1966 zbl = {0141.00801} } +@Article{ KoukoulopoulosMaynard2020, + author = {Koukoulopoulos, Dimitris and Maynard, James}, + title = {On the {D}uffin-{S}chaeffer conjecture}, + journal = {Ann. of Math. (2)}, + fjournal = {Annals of Mathematics. Second Series}, + volume = {192}, + year = {2020}, + number = {1}, + pages = {251--307}, + issn = {0003-486X}, + mrclass = {11J83 (05C40)}, + mrnumber = {4125453}, + mrreviewer = {Sam Chow}, + doi = {10.4007/annals.2020.192.1.5}, + url = {https://doi.org/10.4007/annals.2020.192.1.5}, +} + @Article{ lazarus1973, author = {Michel Lazarus}, title = {Les familles libres maximales d'un module ont-elles le diff --git a/src/analysis/normed/group/basic.lean b/src/analysis/normed/group/basic.lean index fae033dd6ec4c..3eba78c8551d9 100644 --- a/src/analysis/normed/group/basic.lean +++ b/src/analysis/normed/group/basic.lean @@ -1097,6 +1097,16 @@ by simp only [mem_closed_ball, dist_eq_norm_div, mul_div_mul_right_eq_div] a * c ∈ ball (b * c) r ↔ a ∈ ball b r := by simp only [mem_ball, dist_eq_norm_div, mul_div_mul_right_eq_div] +@[to_additive] lemma smul_closed_ball'' : + a • closed_ball b r = closed_ball (a • b) r := +by { ext, simp [mem_closed_ball, set.mem_smul_set, dist_eq_norm_div, div_eq_inv_mul, + ← eq_inv_mul_iff_mul_eq, mul_assoc], } + +@[to_additive] lemma smul_ball'' : + a • ball b r = ball (a • b) r := +by { ext, simp [mem_ball, set.mem_smul_set, dist_eq_norm_div, div_eq_inv_mul, + ← eq_inv_mul_iff_mul_eq, mul_assoc], } + namespace isometric /-- Multiplication `y ↦ x * y` as an `isometry`. -/ diff --git a/src/measure_theory/covering/liminf_limsup.lean b/src/measure_theory/covering/liminf_limsup.lean index bf38500e1ec2d..08303d9914277 100644 --- a/src/measure_theory/covering/liminf_limsup.lean +++ b/src/measure_theory/covering/liminf_limsup.lean @@ -18,6 +18,8 @@ carrying a doubling measure. distances is multiplied by a positive scale factor. This is a generalisation of a result of Cassels, appearing as Lemma 9 on page 217 of [J.W.S. Cassels, *Some metrical theorems in Diophantine approximation. I*](cassels1950). + * `blimsup_thickening_mul_ae_eq`: a variant of `blimsup_cthickening_mul_ae_eq` for thickenings + rather than closed thickenings. -/ @@ -179,6 +181,8 @@ scaled by a positive constant. This lemma is a generalisation of Lemma 9 appearing on page 217 of [J.W.S. Cassels, *Some metrical theorems in Diophantine approximation. I*](cassels1950). +See also `blimsup_thickening_mul_ae_eq`. + NB: The `set : α` type ascription is present because of issue #16932 on GitHub. -/ theorem blimsup_cthickening_mul_ae_eq (p : ℕ → Prop) (s : ℕ → set α) {M : ℝ} (hM : 0 < M) (r : ℕ → ℝ) (hr : tendsto r at_top (𝓝 0)) : @@ -218,3 +222,65 @@ begin blimsup_congr (eventually_of_forall h₂)], exact ae_eq_set_union (this (λ i, p i ∧ 0 < r i) hr') (ae_eq_refl _), end + +lemma blimsup_cthickening_ae_eq_blimsup_thickening + {p : ℕ → Prop} {s : ℕ → set α} {r : ℕ → ℝ} + (hr : tendsto r at_top (𝓝 0)) (hr' : ∀ᶠ i in at_top , p i → 0 < r i) : + (blimsup (λ i, cthickening (r i) (s i)) at_top p : set α) =ᵐ[μ] + (blimsup (λ i, thickening (r i) (s i)) at_top p : set α) := +begin + refine eventually_le_antisymm_iff.mpr ⟨_, has_subset.subset.eventually_le (_ : _ ≤ _)⟩, + { rw eventually_le_congr (blimsup_cthickening_mul_ae_eq μ p s (@one_half_pos ℝ _) r hr).symm + eventually_eq.rfl, + apply has_subset.subset.eventually_le, + change _ ≤ _, + refine mono_blimsup' (hr'.mono $ λ i hi pi, cthickening_subset_thickening' (hi pi) _ (s i)), + nlinarith [hi pi], }, + { exact mono_blimsup (λ i pi, thickening_subset_cthickening _ _), }, +end + +/-- An auxiliary result en route to `blimsup_thickening_mul_ae_eq`. -/ +lemma blimsup_thickening_mul_ae_eq_aux + (p : ℕ → Prop) (s : ℕ → set α) {M : ℝ} (hM : 0 < M) + (r : ℕ → ℝ) (hr : tendsto r at_top (𝓝 0)) (hr' : ∀ᶠ i in at_top , p i → 0 < r i) : + (blimsup (λ i, thickening (M * r i) (s i)) at_top p : set α) =ᵐ[μ] + (blimsup (λ i, thickening (r i) (s i)) at_top p : set α) := +begin + have h₁ := blimsup_cthickening_ae_eq_blimsup_thickening μ hr hr', + have h₂ := blimsup_cthickening_mul_ae_eq μ p s hM r hr, + replace hr : tendsto (λ i, M * r i) at_top (𝓝 0), { convert hr.const_mul M, simp, }, + replace hr' : ∀ᶠ i in at_top , p i → 0 < M * r i := hr'.mono (λ i hi hip, mul_pos hM (hi hip)), + have h₃ := blimsup_cthickening_ae_eq_blimsup_thickening μ hr hr', + exact h₃.symm.trans (h₂.trans h₁), +end + +/-- Given a sequence of subsets `sᵢ` of a metric space, together with a sequence of radii `rᵢ` +such that `rᵢ → 0`, the set of points which belong to infinitely many of the +`rᵢ`-thickenings of `sᵢ` is unchanged almost everywhere for a doubling measure if the `rᵢ` are all +scaled by a positive constant. + +This lemma is a generalisation of Lemma 9 appearing on page 217 of +[J.W.S. Cassels, *Some metrical theorems in Diophantine approximation. I*](cassels1950). + +See also `blimsup_cthickening_mul_ae_eq`. + +NB: The `set : α` type ascription is present because of issue #16932 on GitHub. -/ +theorem blimsup_thickening_mul_ae_eq + (p : ℕ → Prop) (s : ℕ → set α) {M : ℝ} (hM : 0 < M) (r : ℕ → ℝ) (hr : tendsto r at_top (𝓝 0)) : + (blimsup (λ i, thickening (M * r i) (s i)) at_top p : set α) =ᵐ[μ] + (blimsup (λ i, thickening (r i) (s i)) at_top p : set α) := +begin + let q : ℕ → Prop := λ i, p i ∧ 0 < r i, + have h₁ : blimsup (λ i, thickening (r i) (s i)) at_top p = + blimsup (λ i, thickening (r i) (s i)) at_top q, + { refine blimsup_congr' (eventually_of_forall $ λ i h, _), + replace hi : 0 < r i, { contrapose! h, apply thickening_of_nonpos h, }, + simp only [hi, iff_self_and, implies_true_iff], }, + have h₂ : blimsup (λ i, thickening (M * r i) (s i)) at_top p = + blimsup (λ i, thickening (M * r i) (s i)) at_top q, + { refine blimsup_congr' (eventually_of_forall $ λ i h, _), + replace h : 0 < r i, { rw ← zero_lt_mul_left hM, contrapose! h, apply thickening_of_nonpos h, }, + simp only [h, iff_self_and, implies_true_iff], }, + rw [h₁, h₂], + exact blimsup_thickening_mul_ae_eq_aux μ q s hM r hr (eventually_of_forall (λ i hi, hi.2)), +end diff --git a/src/number_theory/well_approximable.lean b/src/number_theory/well_approximable.lean new file mode 100644 index 0000000000000..c3bdd6abc385e --- /dev/null +++ b/src/number_theory/well_approximable.lean @@ -0,0 +1,321 @@ +/- +Copyright (c) 2022 Oliver Nash. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Oliver Nash +-/ +import dynamics.ergodic.add_circle +import measure_theory.covering.liminf_limsup +import data.nat.totient + +/-! +# Well-approximable numbers and Gallagher's ergodic theorem + +Gallagher's ergodic theorem is a result in metric number theory. It thus belongs to that branch of +mathematics concerning arithmetic properties of real numbers which hold almost eveywhere with +respect to the Lebesgue measure. + +Gallagher's theorem concerns the approximation of real numbers by rational numbers. The input is a +sequence of distances `δ₁, δ₂, ...`, and the theorem concerns the set of real numbers `x` for which +there is an infinity of solutions to: +$$ + |x - m/n| < δₙ, +$$ +where the rational number `m/n` is in lowest terms. The result is that for any `δ`, this set is +either almost all `x` or almost no `x`. + +This result was proved by Gallagher in 1959 +[P. Gallagher, *Approximation by reduced fractions*](Gallagher1961). It is formalised here as +`add_circle.add_well_approximable_ae_empty_or_univ` except with `x` belonging to the circle `ℝ ⧸ ℤ` +since this turns out to be more natural. + +Given a particular `δ`, the Duffin-Schaeffer conjecture (now a theorem) gives a criterion for +deciding which of the two cases in the conclusion of Gallagher's theorem actually occurs. It was +proved by Koukoulopoulos and Maynard in 2019 +[D. Koukoulopoulos, J. Maynard, *On the Duffin-Schaeffer conjecture*](KoukoulopoulosMaynard2020). +We do *not* include a formalisation of the Koukoulopoulos-Maynard result here. + +## Main definitions and results: + + * `approx_order_of`: in a seminormed group `A`, given `n : ℕ` and `δ : ℝ`, `approx_order_of A n δ` + is the set of elements within a distance `δ` of a point of order `n`. + * `well_approximable`: in a seminormed group `A`, given a sequence of distances `δ₁, δ₂, ...`, + `well_approximable A δ` is the limsup as `n → ∞` of the sets `approx_order_of A n δₙ`. Thus, it + is the set of points that lie in infinitely many of the sets `approx_order_of A n δₙ`. + * `add_circle.add_well_approximable_ae_empty_or_univ`: *Gallagher's ergodic theorem* says that for + for the (additive) circle `𝕊`, for any sequence of distances `δ`, the set + `add_well_approximable 𝕊 δ` is almost empty or almost full. + +## TODO: + +The hypothesis `hδ` in `add_circle.add_well_approximable_ae_empty_or_univ` can be dropped. +An elementary (non-measure-theoretic) argument shows that if `¬ hδ` holds then +`add_well_approximable 𝕊 δ = univ` (provided `δ` is non-negative). +-/ + +open set filter function metric measure_theory +open_locale measure_theory topological_space pointwise + +/-- In a seminormed group `A`, given `n : ℕ` and `δ : ℝ`, `approx_order_of A n δ` is the set of +elements within a distance `δ` of a point of order `n`. -/ +@[to_additive approx_add_order_of "In a seminormed additive group `A`, given `n : ℕ` and `δ : ℝ`, +`approx_add_order_of A n δ` is the set of elements within a distance `δ` of a point of order `n`."] +def approx_order_of (A : Type*) [seminormed_group A] (n : ℕ) (δ : ℝ) : set A := +thickening δ {y | order_of y = n} + +@[to_additive mem_approx_add_order_of_iff] +lemma mem_approx_order_of_iff {A : Type*} [seminormed_group A] {n : ℕ} {δ : ℝ} {a : A} : + a ∈ approx_order_of A n δ ↔ ∃ (b : A), order_of b = n ∧ a ∈ ball b δ := +by simp only [approx_order_of, thickening_eq_bUnion_ball, mem_Union₂, mem_set_of_eq, exists_prop] + +/-- In a seminormed group `A`, given a sequence of distances `δ₁, δ₂, ...`, `well_approximable A δ` +is the limsup as `n → ∞` of the sets `approx_order_of A n δₙ`. Thus, it is the set of points that +lie in infinitely many of the sets `approx_order_of A n δₙ`. -/ +@[to_additive add_well_approximable "In a seminormed additive group `A`, given a sequence of +distances `δ₁, δ₂, ...`, `add_well_approximable A δ` is the limsup as `n → ∞` of the sets +`approx_add_order_of A n δₙ`. Thus, it is the set of points that lie in infinitely many of the sets +`approx_add_order_of A n δₙ`."] +def well_approximable (A : Type*) [seminormed_group A] (δ : ℕ → ℝ) : set A := +blimsup (λ n, approx_order_of A n (δ n)) at_top (λ n, 0 < n) + +@[to_additive mem_add_well_approximable_iff] +lemma mem_well_approximable_iff {A : Type*} [seminormed_group A] {δ : ℕ → ℝ} {a : A} : + a ∈ well_approximable A δ ↔ a ∈ blimsup (λ n, approx_order_of A n (δ n)) at_top (λ n, 0 < n) := +iff.rfl + +namespace approx_order_of + +variables {A : Type*} [seminormed_comm_group A] {a : A} {m n : ℕ} (δ : ℝ) + +@[to_additive] +lemma image_pow_subset_of_coprime (hm : 0 < m) (hmn : n.coprime m) : + (λ y, y^m) '' (approx_order_of A n δ) ⊆ approx_order_of A n (m * δ) := +begin + rintros - ⟨a, ha, rfl⟩, + obtain ⟨b, hb, hab⟩ := mem_approx_order_of_iff.mp ha, + replace hb : b^m ∈ {u : A | order_of u = n}, { rw ← hb at hmn ⊢, exact order_of_pow_coprime hmn }, + apply ball_subset_thickening hb ((m : ℝ) • δ), + convert pow_mem_ball hm hab using 1, + simp only [nsmul_eq_mul, algebra.id.smul_eq_mul], +end + +@[to_additive] +lemma image_pow_subset (n : ℕ) (hm : 0 < m) : + (λ y, y^m) '' (approx_order_of A (n * m) δ) ⊆ approx_order_of A n (m * δ) := +begin + rintros - ⟨a, ha, rfl⟩, + obtain ⟨b, hb : order_of b = n * m, hab : a ∈ ball b δ⟩ := mem_approx_order_of_iff.mp ha, + replace hb : b^m ∈ {y : A | order_of y = n}, + { rw [mem_set_of_eq, order_of_pow' b hm.ne', hb, nat.gcd_mul_left_left, n.mul_div_cancel hm], }, + apply ball_subset_thickening hb (m * δ), + convert pow_mem_ball hm hab, + simp only [nsmul_eq_mul], +end + +@[to_additive] +lemma smul_subset_of_coprime (han : (order_of a).coprime n) : + a • approx_order_of A n δ ⊆ approx_order_of A ((order_of a) * n) δ := +begin + simp_rw [approx_order_of, thickening_eq_bUnion_ball, ← image_smul, image_Union₂, + image_smul, smul_ball'', smul_eq_mul, mem_set_of_eq], + refine Union₂_subset_iff.mpr (λ b hb c hc, _), + simp only [mem_Union, exists_prop], + refine ⟨a * b, _, hc⟩, + rw ← hb at ⊢ han, + exact (commute.all a b).order_of_mul_eq_mul_order_of_of_coprime han, +end + +@[to_additive vadd_eq_of_mul_dvd] +lemma smul_eq_of_mul_dvd (hn : 0 < n) (han : (order_of a)^2 ∣ n) : + a • approx_order_of A n δ = approx_order_of A n δ := +begin + simp_rw [approx_order_of, thickening_eq_bUnion_ball, ← image_smul, image_Union₂, + image_smul, smul_ball'', smul_eq_mul, mem_set_of_eq], + replace han : ∀ {b : A}, order_of b = n → order_of (a * b) = n, + { intros b hb, + rw ← hb at han hn, + rw sq at han, + rwa [(commute.all a b).order_of_mul_eq_right_of_forall_prime_mul_dvd (order_of_pos_iff.mp hn) + (λ p hp hp', dvd_trans (mul_dvd_mul_right hp' $ order_of a) han)], }, + let f : {b : A | order_of b = n} → {b : A | order_of b = n} := λ b, ⟨a * b, han b.property⟩, + have hf : surjective f, + { rintros ⟨b, hb⟩, + refine ⟨⟨a⁻¹ * b, _⟩, _⟩, + { rw [mem_set_of_eq, ← order_of_inv, mul_inv_rev, inv_inv, mul_comm], + apply han, + simpa, }, + { simp only [subtype.mk_eq_mk, subtype.coe_mk, mul_inv_cancel_left], }, }, + simpa only [f, mem_set_of_eq, subtype.coe_mk, Union_coe_set] using + hf.Union_comp (λ b, ball (b : A) δ), +end + +end approx_order_of + +namespace unit_add_circle + +lemma mem_approx_add_order_of_iff {δ : ℝ} {x : unit_add_circle} {n : ℕ} (hn : 0 < n) : + x ∈ approx_add_order_of unit_add_circle n δ ↔ + ∃ m < n, gcd m n = 1 ∧ ‖x - ↑((m : ℝ) / n)‖ < δ := +begin + haveI : fact ((0 : ℝ) < 1) := ⟨zero_lt_one⟩, + simp only [mem_approx_add_order_of_iff, mem_set_of_eq, ball, exists_prop, dist_eq_norm, + add_circle.add_order_of_eq_pos_iff hn, mul_one], + split, + { rintros ⟨y, ⟨m, hm₁, hm₂, rfl⟩, hx⟩, exact ⟨m, hm₁, hm₂, hx⟩, }, + { rintros ⟨m, hm₁, hm₂, hx⟩, exact ⟨↑((m : ℝ) / n), ⟨m, hm₁, hm₂, rfl⟩, hx⟩, }, +end + +lemma mem_add_well_approximable_iff (δ : ℕ → ℝ) (x : unit_add_circle) : + x ∈ add_well_approximable unit_add_circle δ ↔ + {n : ℕ | ∃ m < n, gcd m n = 1 ∧ ‖x - ↑((m : ℝ) / n)‖ < δ n}.infinite := +begin + simp only [mem_add_well_approximable_iff, ← nat.cofinite_eq_at_top, cofinite.blimsup_set_eq, + mem_set_of_eq], + refine iff_of_eq (congr_arg set.infinite $ ext (λ n, ⟨λ hn, _, λ hn, _⟩)), + { exact (mem_approx_add_order_of_iff hn.1).mp hn.2, }, + { have h : 0 < n := by { obtain ⟨m, hm₁, hm₂, hm₃⟩ := hn, exact pos_of_gt hm₁, }, + exact ⟨h, (mem_approx_add_order_of_iff h).mpr hn⟩, }, +end + +end unit_add_circle + +namespace add_circle + +variables {T : ℝ} [hT : fact (0 < T)] +include hT + +local notation a `∤` b := ¬ a ∣ b +local notation a `∣∣` b := (a ∣ b) ∧ (a*a ∤ b) +local notation `𝕊` := add_circle T + +/-- *Gallagher's ergodic theorem* on Diophantine approximation. -/ +theorem add_well_approximable_ae_empty_or_univ (δ : ℕ → ℝ) (hδ : tendsto δ at_top (𝓝 0)) : + (∀ᵐ x, ¬ add_well_approximable 𝕊 δ x) ∨ ∀ᵐ x, add_well_approximable 𝕊 δ x := +begin + /- Sketch of proof: + + Let `E := add_well_approximable 𝕊 δ`. For each prime `p : ℕ`, we can partition `E` into three + pieces `E = (A p) ∪ (B p) ∪ (C p)` where: + `A p = blimsup (approx_add_order_of 𝕊 n (δ n)) at_top (λ n, 0 < n ∧ (p ∤ n))` + `B p = blimsup (approx_add_order_of 𝕊 n (δ n)) at_top (λ n, 0 < n ∧ (p ∣∣ n))` + `C p = blimsup (approx_add_order_of 𝕊 n (δ n)) at_top (λ n, 0 < n ∧ (p*p ∣ n))`. + (In other words, `A p` is the set of points `x` for which there exist infinitely-many `n` such + that `x` is within a distance `δ n` of a point of order `n` and `p ∤ n`. Similarly for `B`, `C`.) + + These sets have the following key properties: + 1. `A p` is almost invariant under the ergodic map `y ↦ p • y` + 2. `B p` is almost invariant under the ergodic map `y ↦ p • y + 1/p` + 3. `C p` is invariant under the map `y ↦ y + 1/p` + To prove 1 and 2 we need the key result `blimsup_thickening_mul_ae_eq` but 3 is elementary. + + It follows from `add_circle.ergodic_nsmul_add` and `ergodic.ae_empty_or_univ_of_image_ae_le` that + if either `A p` or `B p` is not almost empty for any `p`, then it is almost full and thus so is + `E`. We may therefore assume that both `A p` and `B p` are almost empty for all `p`. We thus have + `E` is almost equal to `C p` for every prime. Combining this with 3 we find that `E` is almost + invariant under the map `y ↦ y + 1/p` for every prime `p`. The required result then follows from + `add_circle.ae_empty_or_univ_of_forall_vadd_ae_eq_self`. -/ + set μ : measure 𝕊 := volume, + set primes : set ℕ := {p | p.prime}, + haveI : nonempty primes := ⟨⟨2, nat.prime_two⟩⟩, + set u : primes → 𝕊 := λ p, ↑(((↑(1 : ℕ) : ℝ) / p) * T), + have hu₀ : ∀ (p : primes), add_order_of (u p) = (p : ℕ), + { rintros ⟨p, hp⟩, exact add_order_of_div_of_gcd_eq_one hp.pos (gcd_one_left p), }, + have hu : tendsto (add_order_of ∘ u) at_top at_top, + { rw (funext hu₀ : add_order_of ∘ u = coe), + have h_mono : monotone (coe : primes → ℕ) := λ p q hpq, hpq, + refine h_mono.tendsto_at_top_at_top (λ n, _), + obtain ⟨p, hp, hp'⟩ := n.exists_infinite_primes, + exact ⟨⟨p, hp'⟩, hp⟩, }, + set E := add_well_approximable 𝕊 δ, + set X : ℕ → set 𝕊 := λ n, approx_add_order_of 𝕊 n (δ n), + set A : ℕ → set 𝕊 := λ p, blimsup X at_top (λ n, 0 < n ∧ (p ∤ n)), + set B : ℕ → set 𝕊 := λ p, blimsup X at_top (λ n, 0 < n ∧ (p ∣∣ n)), + set C : ℕ → set 𝕊 := λ p, blimsup X at_top (λ n, 0 < n ∧ (p^2 ∣ n)), + have hA₀ : ∀ p, measurable_set (A p) := + λ p, measurable_set.measurable_set_blimsup (λ n hn, is_open_thickening.measurable_set), + have hB₀ : ∀ p, measurable_set (B p) := + λ p, measurable_set.measurable_set_blimsup (λ n hn, is_open_thickening.measurable_set), + have hE₀ : null_measurable_set E μ, + { refine (measurable_set.measurable_set_blimsup + (λ n hn, is_open.measurable_set _)).null_measurable_set, + exact is_open_thickening, }, + have hE₁ : ∀ p, E = (A p) ∪ (B p) ∪ (C p), + { intros p, + simp only [E, add_well_approximable, ← blimsup_or_eq_sup, ← and_or_distrib_left, ← sup_eq_union, + sq], + congr, + refine funext (λ n, propext $ iff_self_and.mpr (λ hn, _)), + -- `tauto` can finish from here but unfortunately it's very slow. + simp only [(em (p ∣ n)).symm, (em (p*p ∣ n)).symm, or_and_distrib_left, or_true, true_and, + or_assoc], }, + have hE₂ : ∀ (p : primes), A p =ᵐ[μ] (∅ : set 𝕊) ∧ B p =ᵐ[μ] (∅ : set 𝕊) → E =ᵐ[μ] C p, + { rintros p ⟨hA, hB⟩, + rw hE₁ p, + exact union_ae_eq_right_of_ae_eq_empty ((union_ae_eq_right_of_ae_eq_empty hA).trans hB), }, + have hA : ∀ (p : primes), A p =ᵐ[μ] (∅ : set 𝕊) ∨ A p =ᵐ[μ] univ, + { rintros ⟨p, hp⟩, + let f : 𝕊 → 𝕊 := λ y, (p : ℕ) • y, + suffices : f '' (A p) ⊆ + blimsup (λ n, approx_add_order_of 𝕊 n (p * δ n)) at_top (λ n, 0 < n ∧ (p ∤ n)), + { apply (ergodic_nsmul hp.one_lt).ae_empty_or_univ_of_image_ae_le (hA₀ p), + apply (has_subset.subset.eventually_le this).congr eventually_eq.rfl, + exact blimsup_thickening_mul_ae_eq μ + (λ n, 0 < n ∧ (p ∤ n)) (λ n, {y | add_order_of y = n}) (nat.cast_pos.mpr hp.pos) _ hδ, }, + refine (Sup_hom.set_image f).apply_blimsup_le.trans (mono_blimsup $ λ n hn, _), + replace hn := nat.coprime_comm.mp (hp.coprime_iff_not_dvd.2 hn.2), + exact approx_add_order_of.image_nsmul_subset_of_coprime (δ n) hp.pos hn, }, + have hB : ∀ (p : primes), B p =ᵐ[μ] (∅ : set 𝕊) ∨ B p =ᵐ[μ] univ, + { rintros ⟨p, hp⟩, + let x := u ⟨p, hp⟩, + let f : 𝕊 → 𝕊 := λ y, p • y + x, + suffices : f '' (B p) ⊆ + blimsup (λ n, approx_add_order_of 𝕊 n (p * δ n)) at_top (λ n, 0 < n ∧ (p ∣∣ n)), + { apply (ergodic_nsmul_add x hp.one_lt).ae_empty_or_univ_of_image_ae_le (hB₀ p), + apply (has_subset.subset.eventually_le this).congr eventually_eq.rfl, + exact blimsup_thickening_mul_ae_eq μ + (λ n, 0 < n ∧ (p ∣∣ n)) (λ n, {y | add_order_of y = n}) (nat.cast_pos.mpr hp.pos) _ hδ, }, + refine (Sup_hom.set_image f).apply_blimsup_le.trans (mono_blimsup _), + rintros n ⟨hn, h_div, h_ndiv⟩, + have h_cop : (add_order_of x).coprime (n/p), + { obtain ⟨q, rfl⟩ := h_div, + rw [hu₀, subtype.coe_mk, hp.coprime_iff_not_dvd, q.mul_div_cancel_left hp.pos], + exact λ contra, h_ndiv (mul_dvd_mul_left p contra), }, + replace h_div : n / p * p = n := nat.div_mul_cancel h_div, + have hf : f = (λ y, x + y) ∘ (λ y, p • y), { ext, simp [add_comm x], }, + simp_rw [comp_app], + rw [le_eq_subset, Sup_hom.set_image_to_fun, hf, image_comp], + have := @monotone_image 𝕊 𝕊 (λ y, x + y), + specialize this (approx_add_order_of.image_nsmul_subset (δ n) (n/p) hp.pos), + simp only [h_div] at this ⊢, + refine this.trans _, + convert approx_add_order_of.vadd_subset_of_coprime (p * δ n) h_cop, + simp only [hu₀, subtype.coe_mk, h_div, mul_comm p], }, + change (∀ᵐ x, x ∉ E) ∨ E ∈ volume.ae, + rw [← eventually_eq_empty, ← eventually_eq_univ], + have hC : ∀ (p : primes), (u p) +ᵥ C p = C p, + { intros p, + let e := (add_action.to_perm (u p) : equiv.perm 𝕊).to_order_iso_set, + change e (C p) = C p, + rw [e.apply_blimsup, ← hu₀ p], + exact blimsup_congr (eventually_of_forall $ λ n hn, + approx_add_order_of.vadd_eq_of_mul_dvd (δ n) hn.1 hn.2), }, + by_cases h : ∀ (p : primes), A p =ᵐ[μ] (∅ : set 𝕊) ∧ B p =ᵐ[μ] (∅ : set 𝕊), + { replace h : ∀ (p : primes), ((u p) +ᵥ E : set _) =ᵐ[μ] E, + { intros p, + replace hE₂ : E =ᵐ[μ] C p := hE₂ p (h p), + have h_qmp : measure_theory.measure.quasi_measure_preserving ((+ᵥ) (-u p)) μ μ := + (measure_preserving_vadd _ μ).quasi_measure_preserving, + refine (h_qmp.vadd_ae_eq_of_ae_eq (u p) hE₂).trans (ae_eq_trans _ hE₂.symm), + rw hC, }, + exact ae_empty_or_univ_of_forall_vadd_ae_eq_self hE₀ h hu, }, + { right, + simp only [not_forall, not_and_distrib] at h, + obtain ⟨p, hp⟩ := h, + rw hE₁ p, + cases hp, + { cases hA p, { contradiction, }, + simp only [h, union_ae_eq_univ_of_ae_eq_univ_left], }, + { cases hB p, { contradiction, }, + simp only [h, union_ae_eq_univ_of_ae_eq_univ_left, union_ae_eq_univ_of_ae_eq_univ_right], } }, +end + +end add_circle diff --git a/src/order/liminf_limsup.lean b/src/order/liminf_limsup.lean index effad5afac2bb..8540b8d61e19f 100644 --- a/src/order/liminf_limsup.lean +++ b/src/order/liminf_limsup.lean @@ -544,6 +544,23 @@ theorem has_basis.limsup_eq_infi_supr {p : ι → Prop} {s : ι → set β} {f : (h : f.has_basis p s) : limsup u f = ⨅ i (hi : p i), ⨆ a ∈ s i, u a := (h.map u).Limsup_eq_infi_Sup.trans $ by simp only [Sup_image, id] +lemma blimsup_congr' {f : filter β} {p q : β → Prop} {u : β → α} + (h : ∀ᶠ x in f, u x ≠ ⊥ → (p x ↔ q x)) : + blimsup u f p = blimsup u f q := +begin + simp only [blimsup_eq], + congr, + ext a, + refine eventually_congr (h.mono $ λ b hb, _), + cases eq_or_ne (u b) ⊥ with hu hu, { simp [hu], }, + rw hb hu, +end + +lemma bliminf_congr' {f : filter β} {p q : β → Prop} {u : β → α} + (h : ∀ᶠ x in f, u x ≠ ⊤ → (p x ↔ q x)) : + bliminf u f p = bliminf u f q := +@blimsup_congr' αᵒᵈ β _ _ _ _ _ h + lemma blimsup_eq_infi_bsupr {f : filter β} {p : β → Prop} {u : β → α} : blimsup u f p = ⨅ s ∈ f, ⨆ b (hb : p b ∧ b ∈ s), u b := begin diff --git a/src/topology/metric_space/hausdorff_distance.lean b/src/topology/metric_space/hausdorff_distance.lean index d73ceea17a03b..768bf966e24d7 100644 --- a/src/topology/metric_space/hausdorff_distance.lean +++ b/src/topology/metric_space/hausdorff_distance.lean @@ -921,6 +921,9 @@ end thickening δ ({x} : set X) = ball x δ := by { ext, simp [mem_thickening_iff] } +lemma ball_subset_thickening {x : X} {E : set X} (hx : x ∈ E) (δ : ℝ) : ball x δ ⊆ thickening δ E := +subset.trans (by simp) (thickening_subset_of_subset δ $ singleton_subset_iff.mpr hx) + /-- The (open) `δ`-thickening `thickening δ E` of a subset `E` in a metric space equals the union of balls of radius `δ` centered at points of `E`. -/ lemma thickening_eq_bUnion_ball {δ : ℝ} {E : set X} : From 9aba7801eeecebb61f58a5763c2b6dd1b47dc6ef Mon Sep 17 00:00:00 2001 From: leanprover-community-bot Date: Sun, 1 Jan 2023 04:03:37 +0000 Subject: [PATCH 0121/1291] chore(scripts): update nolints.txt (#18035) I am happy to remove some nolints for you! --- scripts/nolints.txt | 4 ---- 1 file changed, 4 deletions(-) diff --git a/scripts/nolints.txt b/scripts/nolints.txt index b6c360c942572..eb4945fe58bf8 100644 --- a/scripts/nolints.txt +++ b/scripts/nolints.txt @@ -348,10 +348,6 @@ apply_nolint stream.is_bisimulation doc_blame -- geometry/euclidean/circumcenter.lean apply_nolint affine_independent.exists_unique_dist_eq fintype_finite --- group_theory/coset.lean -apply_nolint add_subgroup.card_quotient_dvd_card fintype_finite -apply_nolint subgroup.card_quotient_dvd_card fintype_finite - -- group_theory/group_action/sub_mul_action.lean apply_nolint sub_mul_action.has_zero fails_quickly From f252872231e87a5db80d9938fc05530e70f23a94 Mon Sep 17 00:00:00 2001 From: Chris Hughes Date: Sun, 1 Jan 2023 13:26:36 +0000 Subject: [PATCH 0122/1291] chore(order/closure): improve theorem statement (#18024) --- src/order/closure.lean | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/src/order/closure.lean b/src/order/closure.lean index 347ee87c9174e..d0d83d61be378 100644 --- a/src/order/closure.lean +++ b/src/order/closure.lean @@ -393,14 +393,14 @@ lemma closure_union_closure_subset (x y : α) : l.closure_sup_closure_le x y @[simp] lemma closure_union_closure_left (x y : α) : - (l ((l x) ∪ y) : set β) = l (x ∪ y) := -l.closure_sup_closure_left x y + l ((l x) ∪ y) = l (x ∪ y) := +set_like.coe_injective (l.closure_sup_closure_left x y) @[simp] lemma closure_union_closure_right (x y : α) : l (x ∪ (l y)) = l (x ∪ y) := set_like.coe_injective (l.closure_sup_closure_right x y) -@[simp] lemma closure_union_closure (x y : α) : +lemma closure_union_closure (x y : α) : l ((l x) ∪ (l y)) = l (x ∪ y) := set_like.coe_injective (l.closure_operator.closure_sup_closure x y) From 8a8ac5bbfed5cfd69bd8713e1b774b3a0a67a532 Mon Sep 17 00:00:00 2001 From: Kyle Miller Date: Sun, 1 Jan 2023 15:14:31 +0000 Subject: [PATCH 0123/1291] refactor(data/set/finite): add `finite_to_set` lemmas to simp set (#18026) Making these be `simp` lemmas allows `simp` discharge such simple `set.finite` side-goals. --- src/data/set/finite.lean | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/src/data/set/finite.lean b/src/data/set/finite.lean index 682155ac5babf..641d6920551a0 100644 --- a/src/data/set/finite.lean +++ b/src/data/set/finite.lean @@ -341,7 +341,7 @@ namespace finset /-- Gives a `set.finite` for the `finset` coerced to a `set`. This is a wrapper around `set.to_finite`. -/ -lemma finite_to_set (s : finset α) : (s : set α).finite := set.to_finite _ +@[simp] lemma finite_to_set (s : finset α) : (s : set α).finite := set.to_finite _ @[simp] lemma finite_to_set_to_finset (s : finset α) : s.finite_to_set.to_finset = s := by { ext, rw [set.finite.mem_to_finset, mem_coe] } @@ -350,7 +350,7 @@ end finset namespace multiset -lemma finite_to_set (s : multiset α) : {x | x ∈ s}.finite := +@[simp] lemma finite_to_set (s : multiset α) : {x | x ∈ s}.finite := by { classical, simpa only [← multiset.mem_to_finset] using s.to_finset.finite_to_set } @[simp] lemma finite_to_set_to_finset [decidable_eq α] (s : multiset α) : @@ -359,7 +359,7 @@ by { ext x, simp } end multiset -lemma list.finite_to_set (l : list α) : {x | x ∈ l}.finite := +@[simp] lemma list.finite_to_set (l : list α) : {x | x ∈ l}.finite := (show multiset α, from ⟦l⟧).finite_to_set /-! ### Finite instances From 1e05171a5e8cf18d98d9cf7b207540acb044acae Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Sun, 1 Jan 2023 15:14:32 +0000 Subject: [PATCH 0124/1291] feat(ring_theory/dedekind_domain/ideal): add fractional_ideal.coe_ideal_span_singleton mul_inv and inv_mul lemmas (#18032) --- src/ring_theory/dedekind_domain/ideal.lean | 10 ++++++++++ 1 file changed, 10 insertions(+) diff --git a/src/ring_theory/dedekind_domain/ideal.lean b/src/ring_theory/dedekind_domain/ideal.lean index 36a23c0addf6b..62a116f164c96 100644 --- a/src/ring_theory/dedekind_domain/ideal.lean +++ b/src/ring_theory/dedekind_domain/ideal.lean @@ -137,6 +137,16 @@ open submodule submodule.is_principal (fractional_ideal.span_singleton R₁⁰ x)⁻¹ = fractional_ideal.span_singleton _ (x⁻¹) := fractional_ideal.one_div_span_singleton x +lemma coe_ideal_span_singleton_mul_inv {x : R₁} (hx : x ≠ 0) : + ((ideal.span {x} : ideal R₁) : fractional_ideal R₁⁰ K) * (ideal.span {x} : ideal R₁)⁻¹ = 1 := +by rw [coe_ideal_span_singleton, span_singleton_inv, span_singleton_mul_span_singleton, + mul_inv_cancel $ (map_ne_zero_iff _ $ no_zero_smul_divisors.algebra_map_injective R₁ K).mpr + hx, span_singleton_one] + +lemma coe_ideal_span_singleton_inv_mul {x : R₁} (hx : x ≠ 0) : + ((ideal.span {x} : ideal R₁) : fractional_ideal R₁⁰ K)⁻¹ * (ideal.span {x} : ideal R₁) = 1 := +by rw [mul_comm, coe_ideal_span_singleton_mul_inv K hx] + lemma mul_generator_self_inv {R₁ : Type*} [comm_ring R₁] [algebra R₁ K] [is_localization R₁⁰ K] (I : fractional_ideal R₁⁰ K) [submodule.is_principal (I : submodule R₁ K)] (h : I ≠ 0) : I * fractional_ideal.span_singleton _ (generator (I : submodule R₁ K))⁻¹ = 1 := From 5456d48751c18d7583e6b8e482b45fd6d0de9bc5 Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Mon, 2 Jan 2023 08:34:13 +0000 Subject: [PATCH 0125/1291] feat(ring_theory/ideal/basic): add ideal.span_pair lemmas (#18036) --- src/ring_theory/ideal/basic.lean | 16 ++++++++++++++++ 1 file changed, 16 insertions(+) diff --git a/src/ring_theory/ideal/basic.lean b/src/ring_theory/ideal/basic.lean index b4a03c6b8d823..3eb4be6f0504a 100644 --- a/src/ring_theory/ideal/basic.lean +++ b/src/ring_theory/ideal/basic.lean @@ -260,10 +260,26 @@ begin exact hmax M (lt_of_lt_of_le hPJ hM2) hM1, end +lemma span_pair_comm {x y : α} : (span {x, y} : ideal α) = span {y, x} := +by simp only [span_insert, sup_comm] + theorem mem_span_pair {x y z : α} : z ∈ span ({x, y} : set α) ↔ ∃ a b, a * x + b * y = z := by simp [mem_span_insert, mem_span_singleton', @eq_comm _ _ z] +@[simp] lemma span_pair_add_mul_left {R : Type u} [comm_ring R] {x y : R} (z : R) : + (span {x + y * z, y} : ideal R) = span {x, y} := +begin + ext, + rw [mem_span_pair, mem_span_pair], + exact ⟨λ ⟨a, b, h⟩, ⟨a, b + a * z, by { rw [← h], ring1 }⟩, + λ ⟨a, b, h⟩, ⟨a, b - a * z, by { rw [← h], ring1 }⟩⟩ +end + +@[simp] lemma span_pair_add_mul_right {R : Type u} [comm_ring R] {x y : R} (z : R) : + (span {x, y + x * z} : ideal R) = span {x, y} := +by rw [span_pair_comm, span_pair_add_mul_left, span_pair_comm] + theorem is_maximal.exists_inv {I : ideal α} (hI : I.is_maximal) {x} (hx : x ∉ I) : ∃ y, ∃ i ∈ I, y * x + i = 1 := begin From 448144f7ae193a8990cb7473c9e9a01990f64ac7 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 2 Jan 2023 11:27:23 +0000 Subject: [PATCH 0126/1291] chore(*): remove references to porting PRs (#18031) These no longer reflect the yaml file (which doesn't really track PR numbers any more). Since the next version of the bot will remove them, let's remove them manually first so that the bot diffs are easier to follow. --- src/algebra/abs.lean | 1 - src/algebra/char_zero/defs.lean | 1 - src/algebra/covariant_and_contravariant.lean | 1 - src/algebra/divisibility/basic.lean | 1 - src/algebra/divisibility/units.lean | 1 - src/algebra/euclidean_domain/basic.lean | 1 - src/algebra/euclidean_domain/defs.lean | 1 - src/algebra/field/basic.lean | 1 - src/algebra/field/defs.lean | 1 - src/algebra/group/basic.lean | 1 - src/algebra/group/commutator.lean | 1 - src/algebra/group/commute.lean | 1 - src/algebra/group/defs.lean | 1 - src/algebra/group/ext.lean | 1 - src/algebra/group/inj_surj.lean | 1 - src/algebra/group/opposite.lean | 1 - src/algebra/group/order_synonym.lean | 1 - src/algebra/group/prod.lean | 1 - src/algebra/group/semiconj.lean | 1 - src/algebra/group/type_tags.lean | 1 - src/algebra/group/ulift.lean | 1 - src/algebra/group/units.lean | 1 - src/algebra/group/with_one/basic.lean | 1 - src/algebra/group/with_one/defs.lean | 1 - src/algebra/group/with_one/units.lean | 1 - src/algebra/group_power/basic.lean | 1 - src/algebra/group_power/identities.lean | 1 - src/algebra/group_power/ring.lean | 1 - src/algebra/group_with_zero/basic.lean | 1 - src/algebra/group_with_zero/commute.lean | 1 - src/algebra/group_with_zero/defs.lean | 1 - src/algebra/group_with_zero/divisibility.lean | 1 - src/algebra/group_with_zero/inj_surj.lean | 1 - src/algebra/group_with_zero/semiconj.lean | 1 - src/algebra/group_with_zero/units/basic.lean | 1 - src/algebra/group_with_zero/units/lemmas.lean | 1 - src/algebra/hierarchy_design.lean | 1 - src/algebra/hom/commute.lean | 1 - src/algebra/hom/embedding.lean | 1 - src/algebra/hom/equiv/basic.lean | 1 - src/algebra/hom/equiv/type_tags.lean | 1 - src/algebra/hom/equiv/units/basic.lean | 1 - src/algebra/hom/equiv/units/group_with_zero.lean | 1 - src/algebra/hom/group.lean | 1 - src/algebra/hom/ring.lean | 1 - src/algebra/hom/units.lean | 1 - src/algebra/homology/complex_shape.lean | 1 - src/algebra/invertible.lean | 1 - src/algebra/ne_zero.lean | 1 - src/algebra/opposites.lean | 1 - src/algebra/order/field/canonical/basic.lean | 1 - src/algebra/order/field/canonical/defs.lean | 1 - src/algebra/order/field/defs.lean | 1 - src/algebra/order/field/inj_surj.lean | 1 - src/algebra/order/group/abs.lean | 1 - src/algebra/order/group/defs.lean | 1 - src/algebra/order/group/densely_ordered.lean | 1 - src/algebra/order/group/inj_surj.lean | 1 - src/algebra/order/group/instances.lean | 1 - src/algebra/order/group/min_max.lean | 1 - src/algebra/order/group/order_iso.lean | 1 - src/algebra/order/group/prod.lean | 1 - src/algebra/order/group/type_tags.lean | 3 +-- src/algebra/order/group/units.lean | 1 - src/algebra/order/group/with_top.lean | 1 - src/algebra/order/hom/basic.lean | 1 - src/algebra/order/hom/monoid.lean | 1 - src/algebra/order/monoid/basic.lean | 1 - src/algebra/order/monoid/cancel/basic.lean | 1 - src/algebra/order/monoid/cancel/defs.lean | 1 - src/algebra/order/monoid/canonical/defs.lean | 1 - src/algebra/order/monoid/defs.lean | 1 - src/algebra/order/monoid/lemmas.lean | 1 - src/algebra/order/monoid/min_max.lean | 1 - src/algebra/order/monoid/nat_cast.lean | 1 - src/algebra/order/monoid/order_dual.lean | 3 +-- src/algebra/order/monoid/prod.lean | 3 +-- src/algebra/order/monoid/to_mul_bot.lean | 1 - src/algebra/order/monoid/type_tags.lean | 3 +-- src/algebra/order/monoid/units.lean | 1 - src/algebra/order/monoid/with_top.lean | 3 +-- src/algebra/order/monoid/with_zero/basic.lean | 1 - src/algebra/order/monoid/with_zero/defs.lean | 1 - src/algebra/order/positive/ring.lean | 1 - src/algebra/order/ring/abs.lean | 1 - src/algebra/order/ring/canonical.lean | 1 - src/algebra/order/ring/char_zero.lean | 1 - src/algebra/order/ring/cone.lean | 1 - src/algebra/order/ring/defs.lean | 1 - src/algebra/order/ring/inj_surj.lean | 1 - src/algebra/order/ring/lemmas.lean | 1 - src/algebra/order/sub/basic.lean | 1 - src/algebra/order/sub/canonical.lean | 1 - src/algebra/order/sub/defs.lean | 1 - src/algebra/order/sub/with_top.lean | 1 - src/algebra/order/with_zero.lean | 1 - src/algebra/order/zero_le_one.lean | 1 - src/algebra/pempty_instances.lean | 1 - src/algebra/quotient.lean | 1 - src/algebra/regular/basic.lean | 1 - src/algebra/ring/basic.lean | 1 - src/algebra/ring/commute.lean | 1 - src/algebra/ring/defs.lean | 1 - src/algebra/ring/divisibility.lean | 1 - src/algebra/ring/idempotents.lean | 1 - src/algebra/ring/inj_surj.lean | 1 - src/algebra/ring/order_synonym.lean | 1 - src/algebra/ring/regular.lean | 1 - src/algebra/ring/semiconj.lean | 1 - src/algebra/ring/units.lean | 1 - src/category_theory/category/Rel.lean | 1 - src/category_theory/category/basic.lean | 1 - src/category_theory/concrete_category/bundled.lean | 1 - src/category_theory/functor/basic.lean | 1 - src/category_theory/functor/category.lean | 1 - src/category_theory/functor/functorial.lean | 1 - src/category_theory/isomorphism.lean | 1 - src/category_theory/natural_isomorphism.lean | 1 - src/category_theory/natural_transformation.lean | 1 - src/category_theory/thin.lean | 1 - src/combinatorics/quiver/arborescence.lean | 1 - src/combinatorics/quiver/basic.lean | 1 - src/combinatorics/quiver/cast.lean | 1 - src/combinatorics/quiver/connected_component.lean | 1 - src/combinatorics/quiver/path.lean | 1 - src/combinatorics/quiver/push.lean | 1 - src/combinatorics/quiver/subquiver.lean | 1 - src/control/applicative.lean | 1 - src/control/equiv_functor.lean | 1 - src/control/functor.lean | 1 - src/control/monad/basic.lean | 1 - src/control/traversable/basic.lean | 1 - src/control/traversable/lemmas.lean | 1 - src/control/ulift.lean | 1 - src/data/bool/basic.lean | 1 - src/data/bool/set.lean | 1 - src/data/bracket.lean | 1 - src/data/char.lean | 1 - src/data/countable/defs.lean | 1 - src/data/dlist/basic.lean | 1 - src/data/fin/fin2.lean | 1 - src/data/finite/defs.lean | 1 - src/data/fun_like/basic.lean | 1 - src/data/fun_like/embedding.lean | 1 - src/data/fun_like/equiv.lean | 1 - src/data/int/basic.lean | 1 - src/data/int/cast/basic.lean | 1 - src/data/int/cast/defs.lean | 1 - src/data/int/cast/field.lean | 1 - src/data/int/cast/lemmas.lean | 1 - src/data/int/cast/prod.lean | 1 - src/data/int/div.lean | 1 - src/data/int/dvd/basic.lean | 1 - src/data/int/order/basic.lean | 1 - src/data/int/units.lean | 1 - src/data/lazy_list.lean | 1 - src/data/list/defs.lean | 1 - src/data/list/lex.lean | 1 - src/data/nat/basic.lean | 1 - src/data/nat/cast/basic.lean | 1 - src/data/nat/cast/defs.lean | 1 - src/data/nat/cast/prod.lean | 1 - src/data/nat/cast/with_top.lean | 1 - src/data/nat/gcd/basic.lean | 1 - src/data/nat/order/basic.lean | 1 - src/data/nat/order/lemmas.lean | 1 - src/data/nat/psub.lean | 1 - src/data/nat/set.lean | 1 - src/data/nat/units.lean | 3 +-- src/data/nat/upto.lean | 1 - src/data/num/basic.lean | 1 - src/data/opposite.lean | 1 - src/data/option/basic.lean | 1 - src/data/option/defs.lean | 1 - src/data/option/n_ary.lean | 1 - src/data/pequiv.lean | 1 - src/data/pi/algebra.lean | 1 - src/data/pnat/defs.lean | 1 - src/data/prod/basic.lean | 1 - src/data/prod/lex.lean | 1 - src/data/prod/pprod.lean | 1 - src/data/psigma/order.lean | 1 - src/data/quot.lean | 1 - src/data/rat/init.lean | 1 - src/data/set/basic.lean | 1 - src/data/set/bool_indicator.lean | 1 - src/data/set/image.lean | 1 - src/data/set/n_ary.lean | 1 - src/data/set/opposite.lean | 1 - src/data/set/prod.lean | 1 - src/data/set/sigma.lean | 1 - src/data/set_like/basic.lean | 1 - src/data/sigma/basic.lean | 1 - src/data/sigma/lex.lean | 1 - src/data/sigma/order.lean | 1 - src/data/stream/defs.lean | 1 - src/data/subtype.lean | 1 - src/data/sum/basic.lean | 1 - src/data/sum/order.lean | 1 - src/data/two_pointing.lean | 1 - src/data/ulift.lean | 1 - src/group_theory/eckmann_hilton.lean | 1 - src/group_theory/group_action/defs.lean | 1 - src/group_theory/group_action/option.lean | 1 - src/group_theory/group_action/sigma.lean | 1 - src/group_theory/group_action/sum.lean | 1 - src/group_theory/group_action/units.lean | 1 - src/logic/basic.lean | 1 - src/logic/embedding/basic.lean | 1 - src/logic/embedding/set.lean | 1 - src/logic/equiv/basic.lean | 1 - src/logic/equiv/defs.lean | 1 - src/logic/equiv/embedding.lean | 1 - src/logic/equiv/option.lean | 1 - src/logic/function/basic.lean | 1 - src/logic/function/conjugate.lean | 1 - src/logic/function/iterate.lean | 1 - src/logic/is_empty.lean | 1 - src/logic/lemmas.lean | 1 - src/logic/nonempty.lean | 1 - src/logic/nontrivial.lean | 1 - src/logic/pairwise.lean | 1 - src/logic/relation.lean | 1 - src/logic/relator.lean | 1 - src/logic/unique.lean | 1 - src/order/antisymmetrization.lean | 1 - src/order/basic.lean | 1 - src/order/boolean_algebra.lean | 1 - src/order/bounded_order.lean | 1 - src/order/compare.lean | 1 - src/order/directed.lean | 1 - src/order/disjoint.lean | 1 - src/order/game_add.lean | 1 - src/order/heyting/basic.lean | 1 - src/order/heyting/boundary.lean | 1 - src/order/hom/basic.lean | 1 - src/order/iterate.lean | 1 - src/order/lattice.lean | 1 - src/order/max.lean | 1 - src/order/min_max.lean | 1 - src/order/monotone/basic.lean | 1 - src/order/prop_instances.lean | 1 - src/order/rel_classes.lean | 1 - src/order/rel_iso/basic.lean | 1 - src/order/rel_iso/group.lean | 1 - src/order/rel_iso/set.lean | 1 - src/order/symm_diff.lean | 1 - src/order/synonym.lean | 1 - src/order/well_founded.lean | 1 - src/order/with_bot.lean | 1 - src/ring_theory/coprime/basic.lean | 1 - 251 files changed, 6 insertions(+), 257 deletions(-) diff --git a/src/algebra/abs.lean b/src/algebra/abs.lean index 035d10ec6294d..a7ed47f3b95ca 100644 --- a/src/algebra/abs.lean +++ b/src/algebra/abs.lean @@ -8,7 +8,6 @@ Authors: Christopher Hoskin # Absolute value > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/477 > Any changes to this file require a corresponding PR to mathlib4. This file defines a notational class `has_abs` which adds the unary operator `abs` and the notation diff --git a/src/algebra/char_zero/defs.lean b/src/algebra/char_zero/defs.lean index 6f8fd5beb2081..b816a39338632 100644 --- a/src/algebra/char_zero/defs.lean +++ b/src/algebra/char_zero/defs.lean @@ -9,7 +9,6 @@ import data.int.cast.defs # Characteristic zero > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/661 > Any changes to this file require a corresponding PR to mathlib4. A ring `R` is called of characteristic zero if every natural number `n` is non-zero when considered diff --git a/src/algebra/covariant_and_contravariant.lean b/src/algebra/covariant_and_contravariant.lean index 3825be43ab134..58b4caf68983b 100644 --- a/src/algebra/covariant_and_contravariant.lean +++ b/src/algebra/covariant_and_contravariant.lean @@ -13,7 +13,6 @@ import order.monotone.basic # Covariants and contravariants > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/467 > Any changes to this file require a corresponding PR to mathlib4. This file contains general lemmas and instances to work with the interactions between a relation and diff --git a/src/algebra/divisibility/basic.lean b/src/algebra/divisibility/basic.lean index b374dd07c905b..ffa10bb7ff9c1 100644 --- a/src/algebra/divisibility/basic.lean +++ b/src/algebra/divisibility/basic.lean @@ -11,7 +11,6 @@ import algebra.hom.group # Divisibility > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/833 > Any changes to this file require a corresponding PR to mathlib4. This file defines the basics of the divisibility relation in the context of `(comm_)` `monoid`s. diff --git a/src/algebra/divisibility/units.lean b/src/algebra/divisibility/units.lean index 974e7654eac3f..6cd70984f1a41 100644 --- a/src/algebra/divisibility/units.lean +++ b/src/algebra/divisibility/units.lean @@ -11,7 +11,6 @@ import algebra.group.units # Lemmas about divisibility and units > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/848 > Any changes to this file require a corresponding PR to mathlib4. -/ diff --git a/src/algebra/euclidean_domain/basic.lean b/src/algebra/euclidean_domain/basic.lean index 045ef77015bce..3f485ba27a7f9 100644 --- a/src/algebra/euclidean_domain/basic.lean +++ b/src/algebra/euclidean_domain/basic.lean @@ -13,7 +13,6 @@ import algebra.ring.basic # Lemmas about Euclidean domains > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/919 > Any changes to this file require a corresponding PR to mathlib4. ## Main statements diff --git a/src/algebra/euclidean_domain/defs.lean b/src/algebra/euclidean_domain/defs.lean index cd8945dff425e..5d82d8bc4c87b 100644 --- a/src/algebra/euclidean_domain/defs.lean +++ b/src/algebra/euclidean_domain/defs.lean @@ -12,7 +12,6 @@ import algebra.ring.defs # Euclidean domains > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/871 > Any changes to this file require a corresponding PR to mathlib4. This file introduces Euclidean domains and provides the extended Euclidean algorithm. To be precise, diff --git a/src/algebra/field/basic.lean b/src/algebra/field/basic.lean index 43b82b961b4c0..0beeca085d3fb 100644 --- a/src/algebra/field/basic.lean +++ b/src/algebra/field/basic.lean @@ -12,7 +12,6 @@ import algebra.ring.inj_surj # Lemmas about division (semi)rings and (semi)fields > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/975 > Any changes to this file require a corresponding PR to mathlib4. -/ diff --git a/src/algebra/field/defs.lean b/src/algebra/field/defs.lean index ffbf64f8366df..098aabbec6949 100644 --- a/src/algebra/field/defs.lean +++ b/src/algebra/field/defs.lean @@ -10,7 +10,6 @@ import algebra.ring.defs # Division (semi)rings and (semi)fields > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/668 > Any changes to this file require a corresponding PR to mathlib4. This file introduces fields and division rings (also known as skewfields) and proves some basic diff --git a/src/algebra/group/basic.lean b/src/algebra/group/basic.lean index c8c1c363d72f1..6740b509dae66 100644 --- a/src/algebra/group/basic.lean +++ b/src/algebra/group/basic.lean @@ -10,7 +10,6 @@ import algebra.group.defs # Basic lemmas about semigroups, monoids, and groups > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/457 > Any changes to this file require a corresponding PR to mathlib4. This file lists various basic lemmas about semigroups, monoids, and groups. Most proofs are diff --git a/src/algebra/group/commutator.lean b/src/algebra/group/commutator.lean index 2627a9b8b79fa..874c92a99f0e0 100644 --- a/src/algebra/group/commutator.lean +++ b/src/algebra/group/commutator.lean @@ -11,7 +11,6 @@ import data.bracket # The bracket on a group given by commutator. > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/582 > Any changes to this file require a corresponding PR to mathlib4. -/ diff --git a/src/algebra/group/commute.lean b/src/algebra/group/commute.lean index 34dd712e4b7dc..ede1df9536139 100644 --- a/src/algebra/group/commute.lean +++ b/src/algebra/group/commute.lean @@ -9,7 +9,6 @@ import algebra.group.semiconj # Commuting pairs of elements in monoids > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/750 > Any changes to this file require a corresponding PR to mathlib4. We define the predicate `commute a b := a * b = b * a` and provide some operations on terms `(h : diff --git a/src/algebra/group/defs.lean b/src/algebra/group/defs.lean index 564c7c0eef520..090c13ad34f05 100644 --- a/src/algebra/group/defs.lean +++ b/src/algebra/group/defs.lean @@ -10,7 +10,6 @@ import logic.function.basic # Typeclasses for (semi)groups and monoids > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/457 > Any changes to this file require a corresponding PR to mathlib4. In this file we define typeclasses for algebraic structures with one binary operation. diff --git a/src/algebra/group/ext.lean b/src/algebra/group/ext.lean index f5bccda5bbcf6..94cd7bb95a965 100644 --- a/src/algebra/group/ext.lean +++ b/src/algebra/group/ext.lean @@ -9,7 +9,6 @@ import algebra.hom.group # Extensionality lemmas for monoid and group structures > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/850 > Any changes to this file require a corresponding PR to mathlib4. In this file we prove extensionality lemmas for `monoid` and higher algebraic structures with one diff --git a/src/algebra/group/inj_surj.lean b/src/algebra/group/inj_surj.lean index 8f8f9454388d0..d9c43e3513bdd 100644 --- a/src/algebra/group/inj_surj.lean +++ b/src/algebra/group/inj_surj.lean @@ -11,7 +11,6 @@ import data.int.cast.basic # Lifting algebraic data classes along injective/surjective maps > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/707 > Any changes to this file require a corresponding PR to mathlib4. This file provides definitions that are meant to deal with diff --git a/src/algebra/group/opposite.lean b/src/algebra/group/opposite.lean index 75b4412a2e0c2..72b86f1524c91 100644 --- a/src/algebra/group/opposite.lean +++ b/src/algebra/group/opposite.lean @@ -13,7 +13,6 @@ import data.int.cast.defs # Group structures on the multiplicative and additive opposites > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/912 > Any changes to this file require a corresponding PR to mathlib4. -/ universes u v diff --git a/src/algebra/group/order_synonym.lean b/src/algebra/group/order_synonym.lean index 19fa07d909577..03147174e634e 100644 --- a/src/algebra/group/order_synonym.lean +++ b/src/algebra/group/order_synonym.lean @@ -11,7 +11,6 @@ import order.synonym # Group structure on the order type synonyms > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/651 > Any changes to this file require a corresponding PR to mathlib4. Transfer algebraic instances from `α` to `αᵒᵈ` and `lex α`. diff --git a/src/algebra/group/prod.lean b/src/algebra/group/prod.lean index dbbb9c011d845..3fabcc474d8b4 100644 --- a/src/algebra/group/prod.lean +++ b/src/algebra/group/prod.lean @@ -11,7 +11,6 @@ import algebra.hom.units # Monoid, group etc structures on `M × N` > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/968 > Any changes to this file require a corresponding PR to mathlib4. In this file we define one-binop (`monoid`, `group` etc) structures on `M × N`. We also prove diff --git a/src/algebra/group/semiconj.lean b/src/algebra/group/semiconj.lean index f3d9217d9a060..fc54fd649ee71 100644 --- a/src/algebra/group/semiconj.lean +++ b/src/algebra/group/semiconj.lean @@ -11,7 +11,6 @@ import algebra.group.units # Semiconjugate elements of a semigroup > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/717 > Any changes to this file require a corresponding PR to mathlib4. ## Main definitions diff --git a/src/algebra/group/type_tags.lean b/src/algebra/group/type_tags.lean index 5943f2a83c07e..46f64cf757500 100644 --- a/src/algebra/group/type_tags.lean +++ b/src/algebra/group/type_tags.lean @@ -10,7 +10,6 @@ import data.finite.defs # Type tags that turn additive structures into multiplicative, and vice versa > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/832 > Any changes to this file require a corresponding PR to mathlib4. We define two type tags: diff --git a/src/algebra/group/ulift.lean b/src/algebra/group/ulift.lean index bc65fe96a52a0..49170b7c835cc 100644 --- a/src/algebra/group/ulift.lean +++ b/src/algebra/group/ulift.lean @@ -11,7 +11,6 @@ import algebra.group_with_zero.inj_surj # `ulift` instances for groups and monoids > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/906 > Any changes to this file require a corresponding PR to mathlib4. This file defines instances for group, monoid, semigroup and related structures on `ulift` types. diff --git a/src/algebra/group/units.lean b/src/algebra/group/units.lean index 060aeef30d9fd..67e2461a08917 100644 --- a/src/algebra/group/units.lean +++ b/src/algebra/group/units.lean @@ -11,7 +11,6 @@ import tactic.nontriviality # Units (i.e., invertible elements) of a monoid > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/549 > Any changes to this file require a corresponding PR to mathlib4. An element of a `monoid` is a unit if it has a two-sided inverse. diff --git a/src/algebra/group/with_one/basic.lean b/src/algebra/group/with_one/basic.lean index 7ab93d206bf43..303a93390afda 100644 --- a/src/algebra/group/with_one/basic.lean +++ b/src/algebra/group/with_one/basic.lean @@ -10,7 +10,6 @@ import algebra.hom.equiv.basic # More operations on `with_one` and `with_zero` > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/922 > Any changes to this file require a corresponding PR to mathlib4. This file defines various bundled morphisms on `with_one` and `with_zero` diff --git a/src/algebra/group/with_one/defs.lean b/src/algebra/group/with_one/defs.lean index 202aa8907ec70..813e0abb03ac5 100644 --- a/src/algebra/group/with_one/defs.lean +++ b/src/algebra/group/with_one/defs.lean @@ -10,7 +10,6 @@ import algebra.ring.defs # Adjoining a zero/one to semigroups and related algebraic structures > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/841 > Any changes to this file require a corresponding PR to mathlib4. This file contains different results about adjoining an element to an algebraic structure which then diff --git a/src/algebra/group/with_one/units.lean b/src/algebra/group/with_one/units.lean index 76a2b5e139891..9c28cc1249e65 100644 --- a/src/algebra/group/with_one/units.lean +++ b/src/algebra/group/with_one/units.lean @@ -10,7 +10,6 @@ import algebra.group_with_zero.units.basic # Isomorphism between a group and the units of itself adjoined with `0` > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/955 > Any changes to this file require a corresponding PR to mathlib4. ## Notes diff --git a/src/algebra/group_power/basic.lean b/src/algebra/group_power/basic.lean index 64f57d6890159..281fa0372ef7e 100644 --- a/src/algebra/group_power/basic.lean +++ b/src/algebra/group_power/basic.lean @@ -11,7 +11,6 @@ import algebra.group.type_tags # Power operations on monoids and groups > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/874 > Any changes to this file require a corresponding PR to mathlib4. The power operation on monoids and groups. diff --git a/src/algebra/group_power/identities.lean b/src/algebra/group_power/identities.lean index a5f1ea7d3902b..82982c73a74eb 100644 --- a/src/algebra/group_power/identities.lean +++ b/src/algebra/group_power/identities.lean @@ -9,7 +9,6 @@ import tactic.ring # Identities > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/531 > Any changes to this file require a corresponding PR to mathlib4. This file contains some "named" commutative ring identities. diff --git a/src/algebra/group_power/ring.lean b/src/algebra/group_power/ring.lean index a774790640173..80e3dd7a3d76d 100644 --- a/src/algebra/group_power/ring.lean +++ b/src/algebra/group_power/ring.lean @@ -15,7 +15,6 @@ import data.nat.order.basic # Power operations on monoids with zero, semirings, and rings > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/979 > Any changes to this file require a corresponding PR to mathlib4. This file provides additional lemmas about the natural power operator on rings and semirings. diff --git a/src/algebra/group_with_zero/basic.lean b/src/algebra/group_with_zero/basic.lean index bc47e022764d6..a397f82c1f95a 100644 --- a/src/algebra/group_with_zero/basic.lean +++ b/src/algebra/group_with_zero/basic.lean @@ -11,7 +11,6 @@ import algebra.group.order_synonym # Groups with an adjoined zero element > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/669 > Any changes to this file require a corresponding PR to mathlib4. This file describes structures that are not usually studied on their own right in mathematics, diff --git a/src/algebra/group_with_zero/commute.lean b/src/algebra/group_with_zero/commute.lean index 4ccca1259b5f5..d6fe3e0815dc6 100644 --- a/src/algebra/group_with_zero/commute.lean +++ b/src/algebra/group_with_zero/commute.lean @@ -11,7 +11,6 @@ import tactic.nontriviality # Lemmas about commuting elements in a `monoid_with_zero` or a `group_with_zero`. > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/762 > Any changes to this file require a corresponding PR to mathlib4. -/ diff --git a/src/algebra/group_with_zero/defs.lean b/src/algebra/group_with_zero/defs.lean index 7860f5c1884cc..9cb5b86bc8969 100644 --- a/src/algebra/group_with_zero/defs.lean +++ b/src/algebra/group_with_zero/defs.lean @@ -11,7 +11,6 @@ import algebra.ne_zero # Typeclasses for groups with an adjoined zero element > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/563 > Any changes to this file require a corresponding PR to mathlib4. This file provides just the typeclass definitions, and the projection lemmas that expose their diff --git a/src/algebra/group_with_zero/divisibility.lean b/src/algebra/group_with_zero/divisibility.lean index d9c9c5b1ba160..ae8f2bc07b7ad 100644 --- a/src/algebra/group_with_zero/divisibility.lean +++ b/src/algebra/group_with_zero/divisibility.lean @@ -11,7 +11,6 @@ import algebra.divisibility.units # Divisibility in groups with zero. > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/870 > Any changes to this file require a corresponding PR to mathlib4. Lemmas about divisibility in groups and monoids with zero. diff --git a/src/algebra/group_with_zero/inj_surj.lean b/src/algebra/group_with_zero/inj_surj.lean index d12262651b61d..97b02c66d45d7 100644 --- a/src/algebra/group_with_zero/inj_surj.lean +++ b/src/algebra/group_with_zero/inj_surj.lean @@ -10,7 +10,6 @@ import algebra.group_with_zero.defs # Lifting groups with zero along injective/surjective maps > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/722 > Any changes to this file require a corresponding PR to mathlib4. -/ diff --git a/src/algebra/group_with_zero/semiconj.lean b/src/algebra/group_with_zero/semiconj.lean index cfc9340f2c349..224d3a3204145 100644 --- a/src/algebra/group_with_zero/semiconj.lean +++ b/src/algebra/group_with_zero/semiconj.lean @@ -10,7 +10,6 @@ import algebra.group.semiconj # Lemmas about semiconjugate elements in a `group_with_zero`. > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/757 > Any changes to this file require a corresponding PR to mathlib4. -/ diff --git a/src/algebra/group_with_zero/units/basic.lean b/src/algebra/group_with_zero/units/basic.lean index 83af50744ee66..cc892e9487604 100644 --- a/src/algebra/group_with_zero/units/basic.lean +++ b/src/algebra/group_with_zero/units/basic.lean @@ -12,7 +12,6 @@ import tactic.assert_exists # Lemmas about units in a `monoid_with_zero` or a `group_with_zero`. > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/735 > Any changes to this file require a corresponding PR to mathlib4. We also define `ring.inverse`, a globally defined function on any ring diff --git a/src/algebra/group_with_zero/units/lemmas.lean b/src/algebra/group_with_zero/units/lemmas.lean index 14e3d9e1879b4..68ced6b72bfc9 100644 --- a/src/algebra/group_with_zero/units/lemmas.lean +++ b/src/algebra/group_with_zero/units/lemmas.lean @@ -11,7 +11,6 @@ import group_theory.group_action.units # Further lemmas about units in a `monoid_with_zero` or a `group_with_zero`. > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/920 > Any changes to this file require a corresponding PR to mathlib4. -/ diff --git a/src/algebra/hierarchy_design.lean b/src/algebra/hierarchy_design.lean index 7ba81405e8a93..31c4205bebe12 100644 --- a/src/algebra/hierarchy_design.lean +++ b/src/algebra/hierarchy_design.lean @@ -9,7 +9,6 @@ import tactic.doc_commands # Documentation of the algebraic hierarchy > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/657 > Any changes to this file require a corresponding PR to mathlib4. A library note giving advice on modifying the algebraic hierarchy. diff --git a/src/algebra/hom/commute.lean b/src/algebra/hom/commute.lean index 7622db4ae5c55..24e71649ed215 100644 --- a/src/algebra/hom/commute.lean +++ b/src/algebra/hom/commute.lean @@ -11,7 +11,6 @@ import algebra.group.commute # Multiplicative homomorphisms respect semiconjugation and commutation. > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/831 > Any changes to this file require a corresponding PR to mathlib4. -/ diff --git a/src/algebra/hom/embedding.lean b/src/algebra/hom/embedding.lean index 978feccc9e497..77d654ca61879 100644 --- a/src/algebra/hom/embedding.lean +++ b/src/algebra/hom/embedding.lean @@ -10,7 +10,6 @@ import logic.embedding.basic # The embedding of a cancellative semigroup into itself by multiplication by a fixed element. > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/764 > Any changes to this file require a corresponding PR to mathlib4. -/ diff --git a/src/algebra/hom/equiv/basic.lean b/src/algebra/hom/equiv/basic.lean index 153a5ec737989..423615c762cc7 100644 --- a/src/algebra/hom/equiv/basic.lean +++ b/src/algebra/hom/equiv/basic.lean @@ -12,7 +12,6 @@ import data.pi.algebra # Multiplicative and additive equivs > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/835 > Any changes to this file require a corresponding PR to mathlib4. In this file we define two extensions of `equiv` called `add_equiv` and `mul_equiv`, which are diff --git a/src/algebra/hom/equiv/type_tags.lean b/src/algebra/hom/equiv/type_tags.lean index 329c5a2543094..e08dd40f9a4fe 100644 --- a/src/algebra/hom/equiv/type_tags.lean +++ b/src/algebra/hom/equiv/type_tags.lean @@ -10,7 +10,6 @@ import algebra.group.type_tags # Additive and multiplicative equivalences associated to `multiplicative` and `additive`. > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/943 > Any changes to this file require a corresponding PR to mathlib4. -/ diff --git a/src/algebra/hom/equiv/units/basic.lean b/src/algebra/hom/equiv/units/basic.lean index 1b60a7678fac3..0b25b357a3777 100644 --- a/src/algebra/hom/equiv/units/basic.lean +++ b/src/algebra/hom/equiv/units/basic.lean @@ -10,7 +10,6 @@ import algebra.hom.units # Multiplicative and additive equivalence acting on units. > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/895 > Any changes to this file require a corresponding PR to mathlib4. -/ diff --git a/src/algebra/hom/equiv/units/group_with_zero.lean b/src/algebra/hom/equiv/units/group_with_zero.lean index 5bb1b951fe459..e0144e693d766 100644 --- a/src/algebra/hom/equiv/units/group_with_zero.lean +++ b/src/algebra/hom/equiv/units/group_with_zero.lean @@ -10,7 +10,6 @@ import algebra.group_with_zero.units.basic # Multiplication by a nonzero element in a `group_with_zero` is a permutation. > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/901 > Any changes to this file require a corresponding PR to mathlib4. -/ diff --git a/src/algebra/hom/group.lean b/src/algebra/hom/group.lean index dac0966e0b3d6..17cd97042d8a5 100644 --- a/src/algebra/hom/group.lean +++ b/src/algebra/hom/group.lean @@ -13,7 +13,6 @@ import data.fun_like.basic # Monoid and group homomorphisms > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/659 > Any changes to this file require a corresponding PR to mathlib4. This file defines the bundled structures for monoid and group homomorphisms. Namely, we define diff --git a/src/algebra/hom/ring.lean b/src/algebra/hom/ring.lean index f2d7b2be544ff..5d085245f67a2 100644 --- a/src/algebra/hom/ring.lean +++ b/src/algebra/hom/ring.lean @@ -14,7 +14,6 @@ import data.set.image # Homomorphisms of semirings and rings > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/958 > Any changes to this file require a corresponding PR to mathlib4. This file defines bundled homomorphisms of (non-unital) semirings and rings. As with monoid and diff --git a/src/algebra/hom/units.lean b/src/algebra/hom/units.lean index 0972bdf866549..71f0f7f067f89 100644 --- a/src/algebra/hom/units.lean +++ b/src/algebra/hom/units.lean @@ -9,7 +9,6 @@ import algebra.group.units # Monoid homomorphisms and units > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/745 > Any changes to this file require a corresponding PR to mathlib4. This file allows to lift monoid homomorphisms to group homomorphisms of their units subgroups. It diff --git a/src/algebra/homology/complex_shape.lean b/src/algebra/homology/complex_shape.lean index bffbd9a898316..94b6110670cc5 100644 --- a/src/algebra/homology/complex_shape.lean +++ b/src/algebra/homology/complex_shape.lean @@ -10,7 +10,6 @@ import logic.relation # Shapes of homological complexes > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/635 > Any changes to this file require a corresponding PR to mathlib4. We define a structure `complex_shape ι` for describing the shapes of homological complexes diff --git a/src/algebra/invertible.lean b/src/algebra/invertible.lean index 606c3401e78e5..6728fe56ad325 100644 --- a/src/algebra/invertible.lean +++ b/src/algebra/invertible.lean @@ -12,7 +12,6 @@ import algebra.ring.defs # Invertible elements > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/930 > Any changes to this file require a corresponding PR to mathlib4. This file defines a typeclass `invertible a` for elements `a` with a two-sided diff --git a/src/algebra/ne_zero.lean b/src/algebra/ne_zero.lean index 524afd30716f0..f2bb34bca6feb 100644 --- a/src/algebra/ne_zero.lean +++ b/src/algebra/ne_zero.lean @@ -9,7 +9,6 @@ import logic.basic # `ne_zero` typeclass > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/557 > Any changes to this file require a corresponding PR to mathlib4. We create a typeclass `ne_zero n` which carries around the fact that `(n : R) ≠ 0`. diff --git a/src/algebra/opposites.lean b/src/algebra/opposites.lean index cbe829d2c6471..b3b6f1ab73b4a 100644 --- a/src/algebra/opposites.lean +++ b/src/algebra/opposites.lean @@ -11,7 +11,6 @@ import logic.nontrivial # Multiplicative opposite and algebraic operations on it > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/644 > Any changes to this file require a corresponding PR to mathlib4. In this file we define `mul_opposite α = αᵐᵒᵖ` to be the multiplicative opposite of `α`. It inherits diff --git a/src/algebra/order/field/canonical/basic.lean b/src/algebra/order/field/canonical/basic.lean index 57dd84c1459f7..90e5f56e0bb7d 100644 --- a/src/algebra/order/field/canonical/basic.lean +++ b/src/algebra/order/field/canonical/basic.lean @@ -9,7 +9,6 @@ import algebra.order.field.canonical.defs # Lemmas about canonically ordered semifields. > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/1018 > Any changes to this file require a corresponding PR to mathlib4. -/ diff --git a/src/algebra/order/field/canonical/defs.lean b/src/algebra/order/field/canonical/defs.lean index 674b364673911..30e193fa74c16 100644 --- a/src/algebra/order/field/canonical/defs.lean +++ b/src/algebra/order/field/canonical/defs.lean @@ -11,7 +11,6 @@ import algebra.order.with_zero # Canonically ordered semifields > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/985 > Any changes to this file require a corresponding PR to mathlib4. -/ diff --git a/src/algebra/order/field/defs.lean b/src/algebra/order/field/defs.lean index 1c5bcde47f671..b65b95b79824c 100644 --- a/src/algebra/order/field/defs.lean +++ b/src/algebra/order/field/defs.lean @@ -10,7 +10,6 @@ import algebra.order.ring.defs # Linear ordered (semi)fields > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/905 > Any changes to this file require a corresponding PR to mathlib4. A linear ordered (semi)field is a (semi)field equipped with a linear order such that diff --git a/src/algebra/order/field/inj_surj.lean b/src/algebra/order/field/inj_surj.lean index a9c34901d1726..01a1d0f929096 100644 --- a/src/algebra/order/field/inj_surj.lean +++ b/src/algebra/order/field/inj_surj.lean @@ -11,7 +11,6 @@ import algebra.order.ring.inj_surj # Pulling back linearly ordered fields along injective maps. > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/1017 > Any changes to this file require a corresponding PR to mathlib4. -/ diff --git a/src/algebra/order/group/abs.lean b/src/algebra/order/group/abs.lean index 989793decc56c..da98dd2c757a7 100644 --- a/src/algebra/order/group/abs.lean +++ b/src/algebra/order/group/abs.lean @@ -11,7 +11,6 @@ import order.min_max # Absolute values in ordered groups. > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/896 > Any changes to this file require a corresponding PR to mathlib4. -/ diff --git a/src/algebra/order/group/defs.lean b/src/algebra/order/group/defs.lean index 66a8b7df0b616..69d9fe8ae12a7 100644 --- a/src/algebra/order/group/defs.lean +++ b/src/algebra/order/group/defs.lean @@ -11,7 +11,6 @@ import algebra.order.monoid.cancel.defs # Ordered groups > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/869 > Any changes to this file require a corresponding PR to mathlib4. This file develops the basics of ordered groups. diff --git a/src/algebra/order/group/densely_ordered.lean b/src/algebra/order/group/densely_ordered.lean index 2d2bb5f7cd002..eb80253c69138 100644 --- a/src/algebra/order/group/densely_ordered.lean +++ b/src/algebra/order/group/densely_ordered.lean @@ -11,7 +11,6 @@ import algebra.order.monoid.order_dual # Lemmas about densely linearly ordered groups. > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/956 > Any changes to this file require a corresponding PR to mathlib4. -/ diff --git a/src/algebra/order/group/inj_surj.lean b/src/algebra/order/group/inj_surj.lean index 78bfe0d1a91ce..e26c1fe0f031d 100644 --- a/src/algebra/order/group/inj_surj.lean +++ b/src/algebra/order/group/inj_surj.lean @@ -11,7 +11,6 @@ import algebra.order.group.instances # Pull back ordered groups along injective maps. > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/916 > Any changes to this file require a corresponding PR to mathlib4. -/ diff --git a/src/algebra/order/group/instances.lean b/src/algebra/order/group/instances.lean index c352f5cb860ff..71f3c24131163 100644 --- a/src/algebra/order/group/instances.lean +++ b/src/algebra/order/group/instances.lean @@ -10,7 +10,6 @@ import algebra.order.monoid.order_dual # Additional instances for ordered commutative groups. > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/890 > Any changes to this file require a corresponding PR to mathlib4. -/ diff --git a/src/algebra/order/group/min_max.lean b/src/algebra/order/group/min_max.lean index 9f0cd8d5c1d6a..b0bfc080a6f0a 100644 --- a/src/algebra/order/group/min_max.lean +++ b/src/algebra/order/group/min_max.lean @@ -10,7 +10,6 @@ import algebra.order.monoid.min_max # `min` and `max` in linearly ordered groups. > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/933 > Any changes to this file require a corresponding PR to mathlib4. -/ diff --git a/src/algebra/order/group/order_iso.lean b/src/algebra/order/group/order_iso.lean index 2b068637fe8e3..eadcf1a5b45a7 100644 --- a/src/algebra/order/group/order_iso.lean +++ b/src/algebra/order/group/order_iso.lean @@ -10,7 +10,6 @@ import algebra.hom.equiv.units.basic # Inverse and multiplication as order isomorphisms in ordered groups > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/895 > Any changes to this file require a corresponding PR to mathlib4. -/ diff --git a/src/algebra/order/group/prod.lean b/src/algebra/order/group/prod.lean index 1b0639868d086..3ca4fe7d767ee 100644 --- a/src/algebra/order/group/prod.lean +++ b/src/algebra/order/group/prod.lean @@ -10,7 +10,6 @@ import algebra.order.monoid.prod # Products of ordered commutative groups. > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/1026 > Any changes to this file require a corresponding PR to mathlib4. -/ diff --git a/src/algebra/order/group/type_tags.lean b/src/algebra/order/group/type_tags.lean index d6f1f3776fe1f..96fdf16b6f4df 100644 --- a/src/algebra/order/group/type_tags.lean +++ b/src/algebra/order/group/type_tags.lean @@ -6,10 +6,9 @@ Authors: Jeremy Avigad, Leonardo de Moura, Mario Carneiro, Johannes Hölzl import algebra.order.group.instances import algebra.order.monoid.type_tags -/-! # Ordered group structures on `multiplicative α` and `additive α`. +/-! # Ordered group structures on `multiplicative α` and `additive α`. > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/921 > Any changes to this file require a corresponding PR to mathlib4.-/ variables {α : Type*} diff --git a/src/algebra/order/group/units.lean b/src/algebra/order/group/units.lean index cefbbf4427f22..1a6ae734291e3 100644 --- a/src/algebra/order/group/units.lean +++ b/src/algebra/order/group/units.lean @@ -11,7 +11,6 @@ import algebra.order.monoid.units # Adjoining a top element to a `linear_ordered_add_comm_group_with_top`. > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/898 > Any changes to this file require a corresponding PR to mathlib4. -/ diff --git a/src/algebra/order/group/with_top.lean b/src/algebra/order/group/with_top.lean index 6f23fc042b36b..4a95ca79bbc15 100644 --- a/src/algebra/order/group/with_top.lean +++ b/src/algebra/order/group/with_top.lean @@ -10,7 +10,6 @@ import algebra.order.monoid.with_top # Adjoining a top element to a `linear_ordered_add_comm_group_with_top`. > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/946 > Any changes to this file require a corresponding PR to mathlib4. -/ diff --git a/src/algebra/order/hom/basic.lean b/src/algebra/order/hom/basic.lean index 4f29cd6027e46..bc790b97dd7d9 100644 --- a/src/algebra/order/hom/basic.lean +++ b/src/algebra/order/hom/basic.lean @@ -9,7 +9,6 @@ import algebra.hom.group # Algebraic order homomorphism classes > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/627 > Any changes to this file require a corresponding PR to mathlib4. This file defines hom classes for common properties at the intersection of order theory and algebra. diff --git a/src/algebra/order/hom/monoid.lean b/src/algebra/order/hom/monoid.lean index c3e866f6e61b0..9d36f82d867ce 100644 --- a/src/algebra/order/hom/monoid.lean +++ b/src/algebra/order/hom/monoid.lean @@ -13,7 +13,6 @@ import order.hom.basic # Ordered monoid and group homomorphisms > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/944 > Any changes to this file require a corresponding PR to mathlib4. This file defines morphisms between (additive) ordered monoids. diff --git a/src/algebra/order/monoid/basic.lean b/src/algebra/order/monoid/basic.lean index 05fa36ffe746c..f2a9adb01108a 100644 --- a/src/algebra/order/monoid/basic.lean +++ b/src/algebra/order/monoid/basic.lean @@ -11,7 +11,6 @@ import order.hom.basic # Ordered monoids > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/872 > Any changes to this file require a corresponding PR to mathlib4. This file develops some additional material on ordered monoids. diff --git a/src/algebra/order/monoid/cancel/basic.lean b/src/algebra/order/monoid/cancel/basic.lean index f3d8090e97b84..27fc04ea1e836 100644 --- a/src/algebra/order/monoid/cancel/basic.lean +++ b/src/algebra/order/monoid/cancel/basic.lean @@ -10,7 +10,6 @@ import algebra.order.monoid.cancel.defs # Basic results on ordered cancellative monoids. > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/883 > Any changes to this file require a corresponding PR to mathlib4. We pull back ordered cancellative monoids along injective maps. diff --git a/src/algebra/order/monoid/cancel/defs.lean b/src/algebra/order/monoid/cancel/defs.lean index 84ff915afcc98..d6c15af4ced82 100644 --- a/src/algebra/order/monoid/cancel/defs.lean +++ b/src/algebra/order/monoid/cancel/defs.lean @@ -9,7 +9,6 @@ import algebra.order.monoid.defs # Ordered cancellative monoids > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/774 > Any changes to this file require a corresponding PR to mathlib4. -/ diff --git a/src/algebra/order/monoid/canonical/defs.lean b/src/algebra/order/monoid/canonical/defs.lean index 752f4f8ec5cc8..ffe6d64547c33 100644 --- a/src/algebra/order/monoid/canonical/defs.lean +++ b/src/algebra/order/monoid/canonical/defs.lean @@ -12,7 +12,6 @@ import algebra.order.monoid.defs # Canonically ordered monoids > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/778 > Any changes to this file require a corresponding PR to mathlib4. -/ diff --git a/src/algebra/order/monoid/defs.lean b/src/algebra/order/monoid/defs.lean index 9bd9a2cb423ff..d4052dac4ecbc 100644 --- a/src/algebra/order/monoid/defs.lean +++ b/src/algebra/order/monoid/defs.lean @@ -10,7 +10,6 @@ import order.bounded_order # Ordered monoids > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/771 > Any changes to this file require a corresponding PR to mathlib4. This file provides the definitions of ordered monoids. diff --git a/src/algebra/order/monoid/lemmas.lean b/src/algebra/order/monoid/lemmas.lean index 03107416821fb..2f2ca1f8b5c17 100644 --- a/src/algebra/order/monoid/lemmas.lean +++ b/src/algebra/order/monoid/lemmas.lean @@ -11,7 +11,6 @@ import order.min_max # Ordered monoids > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/608 > Any changes to this file require a corresponding PR to mathlib4. This file develops the basics of ordered monoids. diff --git a/src/algebra/order/monoid/min_max.lean b/src/algebra/order/monoid/min_max.lean index dc98f8e681c75..eb97ace732b46 100644 --- a/src/algebra/order/monoid/min_max.lean +++ b/src/algebra/order/monoid/min_max.lean @@ -10,7 +10,6 @@ import algebra.order.monoid.lemmas # Lemmas about `min` and `max` in an ordered monoid. > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/763 > Any changes to this file require a corresponding PR to mathlib4. -/ diff --git a/src/algebra/order/monoid/nat_cast.lean b/src/algebra/order/monoid/nat_cast.lean index 66f8238dec857..9960b3c1b13e2 100644 --- a/src/algebra/order/monoid/nat_cast.lean +++ b/src/algebra/order/monoid/nat_cast.lean @@ -11,7 +11,6 @@ import data.nat.cast.defs # Order of numerials in an `add_monoid_with_one`. > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/893 > Any changes to this file require a corresponding PR to mathlib4. -/ diff --git a/src/algebra/order/monoid/order_dual.lean b/src/algebra/order/monoid/order_dual.lean index fc0b4f8defdde..dd806089d4316 100644 --- a/src/algebra/order/monoid/order_dual.lean +++ b/src/algebra/order/monoid/order_dual.lean @@ -7,10 +7,9 @@ Authors: Jeremy Avigad, Leonardo de Moura, Mario Carneiro, Johannes Hölzl import algebra.group.order_synonym import algebra.order.monoid.cancel.defs -/-! # Ordered monoid structures on the order dual. +/-! # Ordered monoid structures on the order dual. > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/786 > Any changes to this file require a corresponding PR to mathlib4.-/ universes u diff --git a/src/algebra/order/monoid/prod.lean b/src/algebra/order/monoid/prod.lean index e632cf805920a..b071d078d3dd2 100644 --- a/src/algebra/order/monoid/prod.lean +++ b/src/algebra/order/monoid/prod.lean @@ -7,10 +7,9 @@ import algebra.group.prod import algebra.order.monoid.cancel.defs import algebra.order.monoid.canonical.defs -/-! # Products of ordered monoids +/-! # Products of ordered monoids > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/987 > Any changes to this file require a corresponding PR to mathlib4.-/ namespace prod diff --git a/src/algebra/order/monoid/to_mul_bot.lean b/src/algebra/order/monoid/to_mul_bot.lean index e1f8ea61f6e7c..2cfbed9a89e21 100644 --- a/src/algebra/order/monoid/to_mul_bot.lean +++ b/src/algebra/order/monoid/to_mul_bot.lean @@ -9,7 +9,6 @@ import algebra.order.monoid.type_tags /-! > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/1024 > Any changes to this file require a corresponding PR to mathlib4. Making an additive monoid multiplicative then adding a zero is the same as adding a bottom diff --git a/src/algebra/order/monoid/type_tags.lean b/src/algebra/order/monoid/type_tags.lean index 144f216326f77..69ca44b7e7dc1 100644 --- a/src/algebra/order/monoid/type_tags.lean +++ b/src/algebra/order/monoid/type_tags.lean @@ -7,10 +7,9 @@ import algebra.group.type_tags import algebra.order.monoid.cancel.defs import algebra.order.monoid.canonical.defs -/-! # Ordered monoid structures on `multiplicative α` and `additive α`. +/-! # Ordered monoid structures on `multiplicative α` and `additive α`. > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/873 > Any changes to this file require a corresponding PR to mathlib4.-/ universes u diff --git a/src/algebra/order/monoid/units.lean b/src/algebra/order/monoid/units.lean index 2f834c3d51c55..63ad4dd83958c 100644 --- a/src/algebra/order/monoid/units.lean +++ b/src/algebra/order/monoid/units.lean @@ -11,7 +11,6 @@ import algebra.group.units # Units in ordered monoids > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/873 > Any changes to this file require a corresponding PR to mathlib4. -/ diff --git a/src/algebra/order/monoid/with_top.lean b/src/algebra/order/monoid/with_top.lean index 884be82cf244a..c40041e3a9e03 100644 --- a/src/algebra/order/monoid/with_top.lean +++ b/src/algebra/order/monoid/with_top.lean @@ -8,10 +8,9 @@ import algebra.order.monoid.order_dual import algebra.order.monoid.with_zero.basic import data.nat.cast.defs -/-! # Adjoining top/bottom elements to ordered monoids. +/-! # Adjoining top/bottom elements to ordered monoids. > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/902 > Any changes to this file require a corresponding PR to mathlib4.-/ universes u v diff --git a/src/algebra/order/monoid/with_zero/basic.lean b/src/algebra/order/monoid/with_zero/basic.lean index 09f752abf89a3..730b921ca2722 100644 --- a/src/algebra/order/monoid/with_zero/basic.lean +++ b/src/algebra/order/monoid/with_zero/basic.lean @@ -10,7 +10,6 @@ import algebra.group_with_zero.basic # An instance orphaned from `algebra.order.monoid.with_zero.defs` > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/851 > Any changes to this file require a corresponding PR to mathlib4. We put this here to minimise imports: if you can move it back into diff --git a/src/algebra/order/monoid/with_zero/defs.lean b/src/algebra/order/monoid/with_zero/defs.lean index 5abe890007ec8..4a30575a06993 100644 --- a/src/algebra/order/monoid/with_zero/defs.lean +++ b/src/algebra/order/monoid/with_zero/defs.lean @@ -11,7 +11,6 @@ import algebra.order.zero_le_one # Adjoining a zero element to an ordered monoid. > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/851 > Any changes to this file require a corresponding PR to mathlib4. -/ diff --git a/src/algebra/order/positive/ring.lean b/src/algebra/order/positive/ring.lean index 757201061142d..32bb2e3103b14 100644 --- a/src/algebra/order/positive/ring.lean +++ b/src/algebra/order/positive/ring.lean @@ -10,7 +10,6 @@ import algebra.ring.inj_surj # Algebraic structures on the set of positive numbers > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/911 > Any changes to this file require a corresponding PR to mathlib4. In this file we define various instances (`add_semigroup`, `ordered_comm_monoid` etc) on the diff --git a/src/algebra/order/ring/abs.lean b/src/algebra/order/ring/abs.lean index 24f8dca1b4f8c..4dcdeda4b64dd 100644 --- a/src/algebra/order/ring/abs.lean +++ b/src/algebra/order/ring/abs.lean @@ -12,7 +12,6 @@ import algebra.order.group.abs # Absolute values in linear ordered rings. > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/929 > Any changes to this file require a corresponding PR to mathlib4. -/ diff --git a/src/algebra/order/ring/canonical.lean b/src/algebra/order/ring/canonical.lean index ca2016709b7f9..9727b9fcf4a59 100644 --- a/src/algebra/order/ring/canonical.lean +++ b/src/algebra/order/ring/canonical.lean @@ -11,7 +11,6 @@ import group_theory.group_action.defs # Canoncially ordered rings and semirings. > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/905 > Any changes to this file require a corresponding PR to mathlib4. * `canonically_ordered_comm_semiring` diff --git a/src/algebra/order/ring/char_zero.lean b/src/algebra/order/ring/char_zero.lean index b222522e8345f..abff0e877ff3d 100644 --- a/src/algebra/order/ring/char_zero.lean +++ b/src/algebra/order/ring/char_zero.lean @@ -10,7 +10,6 @@ import algebra.order.ring.defs # Strict ordered semiring have characteristic zero > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/905 > Any changes to this file require a corresponding PR to mathlib4. -/ diff --git a/src/algebra/order/ring/cone.lean b/src/algebra/order/ring/cone.lean index 212f254206158..ee0647139a646 100644 --- a/src/algebra/order/ring/cone.lean +++ b/src/algebra/order/ring/cone.lean @@ -9,7 +9,6 @@ import algebra.order.ring.defs # Constructing an ordered ring from a ring with a specified positive cone. > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/935 > Any changes to this file require a corresponding PR to mathlib4. -/ diff --git a/src/algebra/order/ring/defs.lean b/src/algebra/order/ring/defs.lean index 0b534d98e0973..64e8780473a26 100644 --- a/src/algebra/order/ring/defs.lean +++ b/src/algebra/order/ring/defs.lean @@ -20,7 +20,6 @@ import algebra.group.units # Ordered rings and semirings > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/905 > Any changes to this file require a corresponding PR to mathlib4. This file develops the basics of ordered (semi)rings. diff --git a/src/algebra/order/ring/inj_surj.lean b/src/algebra/order/ring/inj_surj.lean index 87cacef222eaf..e69a498a330bd 100644 --- a/src/algebra/order/ring/inj_surj.lean +++ b/src/algebra/order/ring/inj_surj.lean @@ -11,7 +11,6 @@ import algebra.ring.inj_surj # Pulling back ordered rings along injective maps. > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/917 > Any changes to this file require a corresponding PR to mathlib4. -/ diff --git a/src/algebra/order/ring/lemmas.lean b/src/algebra/order/ring/lemmas.lean index 6427b131a1b0a..2b794a53f1b71 100644 --- a/src/algebra/order/ring/lemmas.lean +++ b/src/algebra/order/ring/lemmas.lean @@ -10,7 +10,6 @@ import algebra.group_with_zero.defs # Multiplication by ·positive· elements is monotonic > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/482 > Any changes to this file require a corresponding PR to mathlib4. Let `α` be a type with `<` and `0`. We use the type `{x : α // 0 < x}` of positive elements of `α` diff --git a/src/algebra/order/sub/basic.lean b/src/algebra/order/sub/basic.lean index f4ca27340619c..82627122e80e0 100644 --- a/src/algebra/order/sub/basic.lean +++ b/src/algebra/order/sub/basic.lean @@ -12,7 +12,6 @@ import algebra.order.sub.defs # Additional results about ordered Subtraction > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/936 > Any changes to this file require a corresponding PR to mathlib4. -/ diff --git a/src/algebra/order/sub/canonical.lean b/src/algebra/order/sub/canonical.lean index 28a2c7cc3859e..0be3913982c1b 100644 --- a/src/algebra/order/sub/canonical.lean +++ b/src/algebra/order/sub/canonical.lean @@ -10,7 +10,6 @@ import algebra.order.sub.defs # Lemmas about subtraction in canonically ordered monoids > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/814 > Any changes to this file require a corresponding PR to mathlib4. -/ diff --git a/src/algebra/order/sub/defs.lean b/src/algebra/order/sub/defs.lean index 56595998a3673..53ad887104d00 100644 --- a/src/algebra/order/sub/defs.lean +++ b/src/algebra/order/sub/defs.lean @@ -12,7 +12,6 @@ import order.lattice # Ordered Subtraction > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/732 > Any changes to this file require a corresponding PR to mathlib4. This file proves lemmas relating (truncated) subtraction with an order. We provide a class diff --git a/src/algebra/order/sub/with_top.lean b/src/algebra/order/sub/with_top.lean index 82589c1cfb386..5e755488d2801 100644 --- a/src/algebra/order/sub/with_top.lean +++ b/src/algebra/order/sub/with_top.lean @@ -10,7 +10,6 @@ import algebra.order.monoid.with_top # Lemma about subtraction in ordered monoids with a top element adjoined. > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/932 > Any changes to this file require a corresponding PR to mathlib4. -/ diff --git a/src/algebra/order/with_zero.lean b/src/algebra/order/with_zero.lean index 02133e056250f..41b268e6c30e2 100644 --- a/src/algebra/order/with_zero.lean +++ b/src/algebra/order/with_zero.lean @@ -15,7 +15,6 @@ import algebra.order.monoid.type_tags # Linearly ordered commutative groups and monoids with a zero element adjoined > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/903 > Any changes to this file require a corresponding PR to mathlib4. This file sets up a special class of linearly ordered commutative monoids diff --git a/src/algebra/order/zero_le_one.lean b/src/algebra/order/zero_le_one.lean index 454b2fc25830c..08714be379a77 100644 --- a/src/algebra/order/zero_le_one.lean +++ b/src/algebra/order/zero_le_one.lean @@ -10,7 +10,6 @@ import algebra.ne_zero # Typeclass expressing `0 ≤ 1`. > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/893 > Any changes to this file require a corresponding PR to mathlib4. -/ diff --git a/src/algebra/pempty_instances.lean b/src/algebra/pempty_instances.lean index 4272ad1efbc76..176f818cb4d65 100644 --- a/src/algebra/pempty_instances.lean +++ b/src/algebra/pempty_instances.lean @@ -11,7 +11,6 @@ import tactic.to_additive # Instances on pempty > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/615 > Any changes to this file require a corresponding PR to mathlib4. This file collects facts about algebraic structures on the (universe-polymorphic) empty type, e.g. diff --git a/src/algebra/quotient.lean b/src/algebra/quotient.lean index 3cd044c5b6bf8..73c3e1c6b819a 100644 --- a/src/algebra/quotient.lean +++ b/src/algebra/quotient.lean @@ -9,7 +9,6 @@ import tactic.basic # Algebraic quotients > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/643 > Any changes to this file require a corresponding PR to mathlib4. This file defines notation for algebraic quotients, e.g. quotient groups `G ⧸ H`, diff --git a/src/algebra/regular/basic.lean b/src/algebra/regular/basic.lean index a4c3aaa361cc9..e71fcb96d2eba 100644 --- a/src/algebra/regular/basic.lean +++ b/src/algebra/regular/basic.lean @@ -11,7 +11,6 @@ import algebra.group_with_zero.basic # Regular elements > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/758 > Any changes to this file require a corresponding PR to mathlib4. We introduce left-regular, right-regular and regular elements, along with their `to_additive` diff --git a/src/algebra/ring/basic.lean b/src/algebra/ring/basic.lean index d0514bfd38581..9b60f938c03a9 100644 --- a/src/algebra/ring/basic.lean +++ b/src/algebra/ring/basic.lean @@ -11,7 +11,6 @@ import algebra.opposites # Semirings and rings > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/830 > Any changes to this file require a corresponding PR to mathlib4. This file gives lemmas about semirings, rings and domains. diff --git a/src/algebra/ring/commute.lean b/src/algebra/ring/commute.lean index b7ed8dbaed8f0..9241b24a54e7e 100644 --- a/src/algebra/ring/commute.lean +++ b/src/algebra/ring/commute.lean @@ -11,7 +11,6 @@ import algebra.group.commute # Semirings and rings > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/759 > Any changes to this file require a corresponding PR to mathlib4. This file gives lemmas about semirings, rings and domains. diff --git a/src/algebra/ring/defs.lean b/src/algebra/ring/defs.lean index 2d75aeec601a2..08afb5f438d3d 100644 --- a/src/algebra/ring/defs.lean +++ b/src/algebra/ring/defs.lean @@ -11,7 +11,6 @@ import data.int.cast.defs # Semirings and rings > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/655 > Any changes to this file require a corresponding PR to mathlib4. This file defines semirings, rings and domains. This is analogous to `algebra.group.defs` and diff --git a/src/algebra/ring/divisibility.lean b/src/algebra/ring/divisibility.lean index f8c2857808642..21dcb248c7b1c 100644 --- a/src/algebra/ring/divisibility.lean +++ b/src/algebra/ring/divisibility.lean @@ -10,7 +10,6 @@ import algebra.ring.defs # Lemmas about divisibility in rings > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/864 > Any changes to this file require a corresponding PR to mathlib4. -/ diff --git a/src/algebra/ring/idempotents.lean b/src/algebra/ring/idempotents.lean index c84331792fea4..1fd69ef91052f 100644 --- a/src/algebra/ring/idempotents.lean +++ b/src/algebra/ring/idempotents.lean @@ -12,7 +12,6 @@ import tactic.nth_rewrite # Idempotents > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/918 > Any changes to this file require a corresponding PR to mathlib4. This file defines idempotents for an arbitary multiplication and proves some basic results, diff --git a/src/algebra/ring/inj_surj.lean b/src/algebra/ring/inj_surj.lean index d52b98aed2d53..9a67f6e8b1ee5 100644 --- a/src/algebra/ring/inj_surj.lean +++ b/src/algebra/ring/inj_surj.lean @@ -11,7 +11,6 @@ import algebra.group_with_zero.inj_surj # Pulling back rings along injective maps, and pushing them forward along surjective maps. > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/734 > Any changes to this file require a corresponding PR to mathlib4. -/ universes u v w x diff --git a/src/algebra/ring/order_synonym.lean b/src/algebra/ring/order_synonym.lean index c046cb781ac5e..a10d6672f5c80 100644 --- a/src/algebra/ring/order_synonym.lean +++ b/src/algebra/ring/order_synonym.lean @@ -10,7 +10,6 @@ import algebra.group.order_synonym # Ring structure on the order type synonyms > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/671 > Any changes to this file require a corresponding PR to mathlib4. Transfer algebraic instances from `α` to `αᵒᵈ` and `lex α`. diff --git a/src/algebra/ring/regular.lean b/src/algebra/ring/regular.lean index 7984f9b3f0f13..e34a13b1a68e8 100644 --- a/src/algebra/ring/regular.lean +++ b/src/algebra/ring/regular.lean @@ -10,7 +10,6 @@ import algebra.ring.defs # Lemmas about regular elements in rings. > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/795 > Any changes to this file require a corresponding PR to mathlib4. -/ diff --git a/src/algebra/ring/semiconj.lean b/src/algebra/ring/semiconj.lean index 8cf0538e47361..ec16a58a95e72 100644 --- a/src/algebra/ring/semiconj.lean +++ b/src/algebra/ring/semiconj.lean @@ -10,7 +10,6 @@ import algebra.ring.defs # Semirings and rings > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/751 > Any changes to this file require a corresponding PR to mathlib4. This file gives lemmas about semirings, rings and domains. diff --git a/src/algebra/ring/units.lean b/src/algebra/ring/units.lean index 5005f799cbd52..c80b6157c2000 100644 --- a/src/algebra/ring/units.lean +++ b/src/algebra/ring/units.lean @@ -10,7 +10,6 @@ import algebra.group.units # Units in semirings and rings > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/746 > Any changes to this file require a corresponding PR to mathlib4. -/ universes u v w x diff --git a/src/category_theory/category/Rel.lean b/src/category_theory/category/Rel.lean index 9297d7d7e1cc6..bc75375e4b5d1 100644 --- a/src/category_theory/category/Rel.lean +++ b/src/category_theory/category/Rel.lean @@ -7,7 +7,6 @@ import category_theory.category.basic /-! > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/822 > Any changes to this file require a corresponding PR to mathlib4. The category of types with binary relations as morphisms. diff --git a/src/category_theory/category/basic.lean b/src/category_theory/category/basic.lean index 8c389c5f0c5c9..02f3ac93913f6 100644 --- a/src/category_theory/category/basic.lean +++ b/src/category_theory/category/basic.lean @@ -9,7 +9,6 @@ import combinatorics.quiver.basic # Categories > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/749 > Any changes to this file require a corresponding PR to mathlib4. Defines a category, as a type class parametrised by the type of objects. diff --git a/src/category_theory/concrete_category/bundled.lean b/src/category_theory/concrete_category/bundled.lean index 99404f87943e6..781996dcc4bbb 100644 --- a/src/category_theory/concrete_category/bundled.lean +++ b/src/category_theory/concrete_category/bundled.lean @@ -9,7 +9,6 @@ import tactic.lint # Bundled types > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/514 > Any changes to this file require a corresponding PR to mathlib4. `bundled c` provides a uniform structure for bundling a type equipped with a type class. diff --git a/src/category_theory/functor/basic.lean b/src/category_theory/functor/basic.lean index 9385d84f3bf37..8af552edc4983 100644 --- a/src/category_theory/functor/basic.lean +++ b/src/category_theory/functor/basic.lean @@ -10,7 +10,6 @@ import category_theory.category.basic # Functors > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/749 > Any changes to this file require a corresponding PR to mathlib4. Defines a functor between categories, extending a `prefunctor` between quivers. diff --git a/src/category_theory/functor/category.lean b/src/category_theory/functor/category.lean index f91b56bef38ee..bf73ec5a9f3f1 100644 --- a/src/category_theory/functor/category.lean +++ b/src/category_theory/functor/category.lean @@ -10,7 +10,6 @@ import category_theory.isomorphism # The category of functors and natural transformations between two fixed categories. > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/749 > Any changes to this file require a corresponding PR to mathlib4. We provide the category instance on `C ⥤ D`, with morphisms the natural transformations. diff --git a/src/category_theory/functor/functorial.lean b/src/category_theory/functor/functorial.lean index 263379bb20e0e..7df2a37790e87 100644 --- a/src/category_theory/functor/functorial.lean +++ b/src/category_theory/functor/functorial.lean @@ -9,7 +9,6 @@ import category_theory.functor.basic # Unbundled functors, as a typeclass decorating the object-level function. > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/822 > Any changes to this file require a corresponding PR to mathlib4. -/ diff --git a/src/category_theory/isomorphism.lean b/src/category_theory/isomorphism.lean index fbaca6813a4e6..d319cb58c5762 100644 --- a/src/category_theory/isomorphism.lean +++ b/src/category_theory/isomorphism.lean @@ -9,7 +9,6 @@ import category_theory.functor.basic # Isomorphisms > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/749 > Any changes to this file require a corresponding PR to mathlib4. This file defines isomorphisms between objects of a category. diff --git a/src/category_theory/natural_isomorphism.lean b/src/category_theory/natural_isomorphism.lean index d49a836724102..7b44dc68f968b 100644 --- a/src/category_theory/natural_isomorphism.lean +++ b/src/category_theory/natural_isomorphism.lean @@ -10,7 +10,6 @@ import category_theory.isomorphism # Natural isomorphisms > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/820 > Any changes to this file require a corresponding PR to mathlib4. For the most part, natural isomorphisms are just another sort of isomorphism. diff --git a/src/category_theory/natural_transformation.lean b/src/category_theory/natural_transformation.lean index 524d20db34af0..872826db90422 100644 --- a/src/category_theory/natural_transformation.lean +++ b/src/category_theory/natural_transformation.lean @@ -9,7 +9,6 @@ import category_theory.functor.basic # Natural transformations > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/749 > Any changes to this file require a corresponding PR to mathlib4. Defines natural transformations between functors. diff --git a/src/category_theory/thin.lean b/src/category_theory/thin.lean index 174ec06b38aca..b402d5135f7ae 100644 --- a/src/category_theory/thin.lean +++ b/src/category_theory/thin.lean @@ -10,7 +10,6 @@ import category_theory.isomorphism # Thin categories > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/822 > Any changes to this file require a corresponding PR to mathlib4. A thin category (also known as a sparse category) is a category with at most one morphism between diff --git a/src/combinatorics/quiver/arborescence.lean b/src/combinatorics/quiver/arborescence.lean index 923b7fd149f16..9bc048ce718bd 100644 --- a/src/combinatorics/quiver/arborescence.lean +++ b/src/combinatorics/quiver/arborescence.lean @@ -12,7 +12,6 @@ import combinatorics.quiver.path # Arborescences > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/997 > Any changes to this file require a corresponding PR to mathlib4. A quiver `V` is an arborescence (or directed rooted tree) when we have a root vertex `root : V` such diff --git a/src/combinatorics/quiver/basic.lean b/src/combinatorics/quiver/basic.lean index 797af55d2b432..1dc41a020da10 100644 --- a/src/combinatorics/quiver/basic.lean +++ b/src/combinatorics/quiver/basic.lean @@ -9,7 +9,6 @@ import data.opposite # Quivers > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/749 > Any changes to this file require a corresponding PR to mathlib4. This module defines quivers. A quiver on a type `V` of vertices assigns to every diff --git a/src/combinatorics/quiver/cast.lean b/src/combinatorics/quiver/cast.lean index d4a2de4e3bae9..62b6fad9209ca 100644 --- a/src/combinatorics/quiver/cast.lean +++ b/src/combinatorics/quiver/cast.lean @@ -11,7 +11,6 @@ import combinatorics.quiver.path # Rewriting arrows and paths along vertex equalities > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/990 > Any changes to this file require a corresponding PR to mathlib4. This files defines `hom.cast` and `path.cast` (and associated lemmas) in order to allow diff --git a/src/combinatorics/quiver/connected_component.lean b/src/combinatorics/quiver/connected_component.lean index 1a27252fabbaa..c70c11e891995 100644 --- a/src/combinatorics/quiver/connected_component.lean +++ b/src/combinatorics/quiver/connected_component.lean @@ -10,7 +10,6 @@ import combinatorics.quiver.symmetric ## Weakly connected components > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/836 > Any changes to this file require a corresponding PR to mathlib4. diff --git a/src/combinatorics/quiver/path.lean b/src/combinatorics/quiver/path.lean index 85c2f3ea2ae94..65fbe46496e62 100644 --- a/src/combinatorics/quiver/path.lean +++ b/src/combinatorics/quiver/path.lean @@ -10,7 +10,6 @@ import logic.lemmas # Paths in quivers > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/811 > Any changes to this file require a corresponding PR to mathlib4. Given a quiver `V`, we define the type of paths from `a : V` to `b : V` as an inductive diff --git a/src/combinatorics/quiver/push.lean b/src/combinatorics/quiver/push.lean index 6fc5366b12342..ee82a7c6ecd3e 100644 --- a/src/combinatorics/quiver/push.lean +++ b/src/combinatorics/quiver/push.lean @@ -9,7 +9,6 @@ import combinatorics.quiver.basic # Pushing a quiver structure along a map > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/868 > Any changes to this file require a corresponding PR to mathlib4. Given a map `σ : V → W` and a `quiver` instance on `V`, this files defines a `quiver` instance diff --git a/src/combinatorics/quiver/subquiver.lean b/src/combinatorics/quiver/subquiver.lean index 28065caed1892..c9db6c0715667 100644 --- a/src/combinatorics/quiver/subquiver.lean +++ b/src/combinatorics/quiver/subquiver.lean @@ -10,7 +10,6 @@ import combinatorics.quiver.basic ## Wide subquivers > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/828 > Any changes to this file require a corresponding PR to mathlib4. A wide subquiver `H` of a quiver `H` consists of a subset of the edge set `a ⟶ b` for diff --git a/src/control/applicative.lean b/src/control/applicative.lean index c2f519392b7b8..7afdd96e3367f 100644 --- a/src/control/applicative.lean +++ b/src/control/applicative.lean @@ -10,7 +10,6 @@ import control.functor # `applicative` instances > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/798 > Any changes to this file require a corresponding PR to mathlib4. This file provides `applicative` instances for concrete functors: diff --git a/src/control/equiv_functor.lean b/src/control/equiv_functor.lean index bf80dc9194bb2..b5588efafca50 100644 --- a/src/control/equiv_functor.lean +++ b/src/control/equiv_functor.lean @@ -9,7 +9,6 @@ import logic.equiv.defs # Functions functorial with respect to equivalences > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/649 > Any changes to this file require a corresponding PR to mathlib4. An `equiv_functor` is a function from `Type → Type` equipped with the additional data of diff --git a/src/control/functor.lean b/src/control/functor.lean index b9decb492e97c..c18cd665db28e 100644 --- a/src/control/functor.lean +++ b/src/control/functor.lean @@ -10,7 +10,6 @@ import control.basic # Functors > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/612 > Any changes to this file require a corresponding PR to mathlib4. This module provides additional lemmas, definitions, and instances for `functor`s. diff --git a/src/control/monad/basic.lean b/src/control/monad/basic.lean index 06fdadec99435..916599b5d74d1 100644 --- a/src/control/monad/basic.lean +++ b/src/control/monad/basic.lean @@ -10,7 +10,6 @@ import tactic.basic # Monad > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/752 > Any changes to this file require a corresponding PR to mathlib4. ## Attributes diff --git a/src/control/traversable/basic.lean b/src/control/traversable/basic.lean index bdc3ed837a7a4..6f297f03b087b 100644 --- a/src/control/traversable/basic.lean +++ b/src/control/traversable/basic.lean @@ -10,7 +10,6 @@ import tactic.ext # Traversable type class > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/788 > Any changes to this file require a corresponding PR to mathlib4. Type classes for traversing collections. The concepts and laws are taken from diff --git a/src/control/traversable/lemmas.lean b/src/control/traversable/lemmas.lean index 87e4df83e52d4..ec36c888752af 100644 --- a/src/control/traversable/lemmas.lean +++ b/src/control/traversable/lemmas.lean @@ -10,7 +10,6 @@ import control.traversable.basic # Traversing collections > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/948 > Any changes to this file require a corresponding PR to mathlib4. This file proves basic properties of traversable and applicative functors and defines diff --git a/src/control/ulift.lean b/src/control/ulift.lean index 8c228a6ce1890..0570be8666309 100644 --- a/src/control/ulift.lean +++ b/src/control/ulift.lean @@ -8,7 +8,6 @@ Authors: Scott Morrison, Jannis Limperg # Monadic instances for `ulift` and `plift` > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/638 > Any changes to this file require a corresponding PR to mathlib4. In this file we define `monad` and `is_lawful_monad` instances on `plift` and `ulift`. -/ diff --git a/src/data/bool/basic.lean b/src/data/bool/basic.lean index bc94fe07905be..d1a90f2db5395 100644 --- a/src/data/bool/basic.lean +++ b/src/data/bool/basic.lean @@ -8,7 +8,6 @@ Authors: Leonardo de Moura, Jeremy Avigad # booleans > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/534 > Any changes to this file require a corresponding PR to mathlib4. This file proves various trivial lemmas about booleans and their diff --git a/src/data/bool/set.lean b/src/data/bool/set.lean index 8128f19f96d33..87ccdb45690e8 100644 --- a/src/data/bool/set.lean +++ b/src/data/bool/set.lean @@ -10,7 +10,6 @@ import data.set.image # Booleans and set operations > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/960 > Any changes to this file require a corresponding PR to mathlib4. This file contains two trivial lemmas about `bool`, `set.univ`, and `set.range`. diff --git a/src/data/bracket.lean b/src/data/bracket.lean index f081f02f9187e..c0b544c455aed 100644 --- a/src/data/bracket.lean +++ b/src/data/bracket.lean @@ -8,7 +8,6 @@ Authors: Patrick Lutz, Oliver Nash # Bracket Notation > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/480 > Any changes to this file require a corresponding PR to mathlib4. This file provides notation which can be used for the Lie bracket, for the commutator of two diff --git a/src/data/char.lean b/src/data/char.lean index 26f4230002fec..886537912db32 100644 --- a/src/data/char.lean +++ b/src/data/char.lean @@ -8,7 +8,6 @@ Authors: Mario Carneiro # More `char` instances > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/609 > Any changes to this file require a corresponding PR to mathlib4. This file provides a `linear_order` instance on `char`. `char` is the type of Unicode scalar values. diff --git a/src/data/countable/defs.lean b/src/data/countable/defs.lean index ecbe0be30c03d..c24a9c97eab72 100644 --- a/src/data/countable/defs.lean +++ b/src/data/countable/defs.lean @@ -9,7 +9,6 @@ import data.finite.defs # Countable types > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/736 > Any changes to this file require a corresponding PR to mathlib4. In this file we define a typeclass saying that a given `Sort*` is countable. See also `encodable` diff --git a/src/data/dlist/basic.lean b/src/data/dlist/basic.lean index 2842d0978711e..d73414a31a5ae 100644 --- a/src/data/dlist/basic.lean +++ b/src/data/dlist/basic.lean @@ -9,7 +9,6 @@ import data.dlist # Difference list > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/616 > Any changes to this file require a corresponding PR to mathlib4. This file provides a few results about `dlist`, which is defined in core Lean. diff --git a/src/data/fin/fin2.lean b/src/data/fin/fin2.lean index ecfc779d73980..97382f66c75f2 100644 --- a/src/data/fin/fin2.lean +++ b/src/data/fin/fin2.lean @@ -8,7 +8,6 @@ Authors: Mario Carneiro # Inductive type variant of `fin` > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/478 > Any changes to this file require a corresponding PR to mathlib4. `fin` is defined as a subtype of `ℕ`. This file defines an equivalent type, `fin2`, which is diff --git a/src/data/finite/defs.lean b/src/data/finite/defs.lean index 8c66af8d3de5f..9c6ecfee5479a 100644 --- a/src/data/finite/defs.lean +++ b/src/data/finite/defs.lean @@ -9,7 +9,6 @@ import logic.equiv.basic # Definition of the `finite` typeclass > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/698 > Any changes to this file require a corresponding PR to mathlib4. This file defines a typeclass `finite` saying that `α : Sort*` is finite. A type is `finite` if it diff --git a/src/data/fun_like/basic.lean b/src/data/fun_like/basic.lean index c250c6d4e5a27..dd8d40155d436 100644 --- a/src/data/fun_like/basic.lean +++ b/src/data/fun_like/basic.lean @@ -12,7 +12,6 @@ import tactic.norm_cast # Typeclass for a type `F` with an injective map to `A → B` > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/541 > Any changes to this file require a corresponding PR to mathlib4. This typeclass is primarily for use by homomorphisms like `monoid_hom` and `linear_map`. diff --git a/src/data/fun_like/embedding.lean b/src/data/fun_like/embedding.lean index e896144a9f438..fa5f28da6027a 100644 --- a/src/data/fun_like/embedding.lean +++ b/src/data/fun_like/embedding.lean @@ -10,7 +10,6 @@ import data.fun_like.basic # Typeclass for a type `F` with an injective map to `A ↪ B` > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/541 > Any changes to this file require a corresponding PR to mathlib4. This typeclass is primarily for use by embeddings such as `rel_embedding`. diff --git a/src/data/fun_like/equiv.lean b/src/data/fun_like/equiv.lean index e6c2dc8d72dbf..d32322e2a0958 100644 --- a/src/data/fun_like/equiv.lean +++ b/src/data/fun_like/equiv.lean @@ -10,7 +10,6 @@ import data.fun_like.embedding # Typeclass for a type `F` with an injective map to `A ≃ B` > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/541 > Any changes to this file require a corresponding PR to mathlib4. This typeclass is primarily for use by isomorphisms like `monoid_equiv` and `linear_equiv`. diff --git a/src/data/int/basic.lean b/src/data/int/basic.lean index 36c01318130ec..52b27c99b41d9 100644 --- a/src/data/int/basic.lean +++ b/src/data/int/basic.lean @@ -10,7 +10,6 @@ import order.monotone.basic # Basic instances on the integers > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/584 > Any changes to this file require a corresponding PR to mathlib4. This file contains: diff --git a/src/data/int/cast/basic.lean b/src/data/int/cast/basic.lean index 913d5db16ef2f..32a9faf7c8004 100644 --- a/src/data/int/cast/basic.lean +++ b/src/data/int/cast/basic.lean @@ -10,7 +10,6 @@ import algebra.group.basic # Cast of integers (additional theorems) > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/670 > Any changes to this file require a corresponding PR to mathlib4. This file proves additional properties about the *canonical* homomorphism from diff --git a/src/data/int/cast/defs.lean b/src/data/int/cast/defs.lean index 2579953147c04..bcb87ad797ead 100644 --- a/src/data/int/cast/defs.lean +++ b/src/data/int/cast/defs.lean @@ -9,7 +9,6 @@ import data.nat.cast.defs # Cast of integers > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/641 > Any changes to this file require a corresponding PR to mathlib4. This file defines the *canonical* homomorphism from the integers into an diff --git a/src/data/int/cast/field.lean b/src/data/int/cast/field.lean index ac4d838b73687..2fe7363a6d4f2 100644 --- a/src/data/int/cast/field.lean +++ b/src/data/int/cast/field.lean @@ -11,7 +11,6 @@ import algebra.group_with_zero.units.lemmas # Cast of integers into fields > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/1016 > Any changes to this file require a corresponding PR to mathlib4. This file concerns the canonical homomorphism `ℤ → F`, where `F` is a field. diff --git a/src/data/int/cast/lemmas.lean b/src/data/int/cast/lemmas.lean index a08b6ceed64ac..6598a6553085a 100644 --- a/src/data/int/cast/lemmas.lean +++ b/src/data/int/cast/lemmas.lean @@ -10,7 +10,6 @@ import data.nat.cast.basic # Cast of integers (additional theorems) > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/995 > Any changes to this file require a corresponding PR to mathlib4. This file proves additional properties about the *canonical* homomorphism from diff --git a/src/data/int/cast/prod.lean b/src/data/int/cast/prod.lean index c4b1897f91a61..ab51bcbacc89d 100644 --- a/src/data/int/cast/prod.lean +++ b/src/data/int/cast/prod.lean @@ -10,7 +10,6 @@ import data.nat.cast.prod # The product of two `add_group_with_one`s. > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/1015 > Any changes to this file require a corresponding PR to mathlib4. -/ diff --git a/src/data/int/div.lean b/src/data/int/div.lean index 4d0724f1449bb..9b44e3a92579a 100644 --- a/src/data/int/div.lean +++ b/src/data/int/div.lean @@ -11,7 +11,6 @@ import algebra.ring.regular # Lemmas relating `/` in `ℤ` with the ordering. > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/1011 > Any changes to this file require a corresponding PR to mathlib4. -/ diff --git a/src/data/int/dvd/basic.lean b/src/data/int/dvd/basic.lean index 24c7c5c92ced2..034d16b7c309a 100644 --- a/src/data/int/dvd/basic.lean +++ b/src/data/int/dvd/basic.lean @@ -10,7 +10,6 @@ import data.nat.cast.basic # Basic lemmas about the divisibility relation in `ℤ`. > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/996 > Any changes to this file require a corresponding PR to mathlib4. -/ diff --git a/src/data/int/order/basic.lean b/src/data/int/order/basic.lean index d91f7a5b7f6d1..de24db37ca9d1 100644 --- a/src/data/int/order/basic.lean +++ b/src/data/int/order/basic.lean @@ -13,7 +13,6 @@ import tactic.assert_exists # Order instances on the integers > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/938 > Any changes to this file require a corresponding PR to mathlib4. This file contains: diff --git a/src/data/int/units.lean b/src/data/int/units.lean index 925c287d972fa..7184fa0163efd 100644 --- a/src/data/int/units.lean +++ b/src/data/int/units.lean @@ -12,7 +12,6 @@ import algebra.ring.units # Lemmas about units in `ℤ`. > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/807 > Any changes to this file require a corresponding PR to mathlib4. -/ diff --git a/src/data/lazy_list.lean b/src/data/lazy_list.lean index b73c4bacd23b3..77c3684dea281 100644 --- a/src/data/lazy_list.lean +++ b/src/data/lazy_list.lean @@ -8,7 +8,6 @@ Authors: Leonardo de Moura # Lazy lists > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/686 > Any changes to this file require a corresponding PR to mathlib4. The type `lazy_list α` is a lazy list with elements of type `α`. diff --git a/src/data/list/defs.lean b/src/data/list/defs.lean index bb582f4881f88..b20ef680b6575 100644 --- a/src/data/list/defs.lean +++ b/src/data/list/defs.lean @@ -11,7 +11,6 @@ import data.rbtree.default_lt ## Definitions on lists > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/803 > Any changes to this file require a corresponding PR to mathlib4. This file contains various definitions on lists. It does not contain diff --git a/src/data/list/lex.lean b/src/data/list/lex.lean index 5b0471b8b4674..3a0c574b37557 100644 --- a/src/data/list/lex.lean +++ b/src/data/list/lex.lean @@ -9,7 +9,6 @@ import order.rel_classes # Lexicographic ordering of lists. > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/672 > Any changes to this file require a corresponding PR to mathlib4. The lexicographic order on `list α` is defined by `L < M` iff diff --git a/src/data/nat/basic.lean b/src/data/nat/basic.lean index 4c4b67c7be4c8..4b5a0b8a04040 100644 --- a/src/data/nat/basic.lean +++ b/src/data/nat/basic.lean @@ -11,7 +11,6 @@ import algebra.ring.defs # Basic operations on the natural numbers > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/729 > Any changes to this file require a corresponding PR to mathlib4. This file contains: diff --git a/src/data/nat/cast/basic.lean b/src/data/nat/cast/basic.lean index 56c3f44f5730a..adee13377c6d0 100644 --- a/src/data/nat/cast/basic.lean +++ b/src/data/nat/cast/basic.lean @@ -15,7 +15,6 @@ import algebra.group.opposite # Cast of natural numbers (additional theorems) > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/980 > Any changes to this file require a corresponding PR to mathlib4. This file proves additional properties about the *canonical* homomorphism from diff --git a/src/data/nat/cast/defs.lean b/src/data/nat/cast/defs.lean index c459e57ab6c57..d508ddac3e2da 100644 --- a/src/data/nat/cast/defs.lean +++ b/src/data/nat/cast/defs.lean @@ -10,7 +10,6 @@ import algebra.ne_zero # Cast of natural numbers > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/641 > Any changes to this file require a corresponding PR to mathlib4. This file defines the *canonical* homomorphism from the natural numbers into an diff --git a/src/data/nat/cast/prod.lean b/src/data/nat/cast/prod.lean index fa8de3dd8438b..22e22e1a188af 100644 --- a/src/data/nat/cast/prod.lean +++ b/src/data/nat/cast/prod.lean @@ -10,7 +10,6 @@ import algebra.group.prod # The product of two `add_monoid_with_one`s. > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/1010 > Any changes to this file require a corresponding PR to mathlib4. -/ diff --git a/src/data/nat/cast/with_top.lean b/src/data/nat/cast/with_top.lean index ddd74be02e52f..f80b27a9a9405 100644 --- a/src/data/nat/cast/with_top.lean +++ b/src/data/nat/cast/with_top.lean @@ -10,7 +10,6 @@ import data.nat.basic # Lemma about the coercion `ℕ → with_bot ℕ`. > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/1019 > Any changes to this file require a corresponding PR to mathlib4. An orphaned lemma about casting from `ℕ` to `with_bot ℕ`, diff --git a/src/data/nat/gcd/basic.lean b/src/data/nat/gcd/basic.lean index fd5e2cf322c9a..c06cb1589bfbd 100644 --- a/src/data/nat/gcd/basic.lean +++ b/src/data/nat/gcd/basic.lean @@ -11,7 +11,6 @@ import data.nat.order.lemmas # Definitions and properties of `nat.gcd`, `nat.lcm`, and `nat.coprime` > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/965 > Any changes to this file require a corresponding PR to mathlib4. Generalizations of these are provided in a later file as `gcd_monoid.gcd` and diff --git a/src/data/nat/order/basic.lean b/src/data/nat/order/basic.lean index 39e36ad9fc3ad..e1a9995b9225f 100644 --- a/src/data/nat/order/basic.lean +++ b/src/data/nat/order/basic.lean @@ -10,7 +10,6 @@ import data.nat.basic # The natural numbers as a linearly ordered commutative semiring > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/907 > Any changes to this file require a corresponding PR to mathlib4. We also have a variety of lemmas which have been deferred from `data.nat.basic` because it is diff --git a/src/data/nat/order/lemmas.lean b/src/data/nat/order/lemmas.lean index 8f244741fd277..29c162696faec 100644 --- a/src/data/nat/order/lemmas.lean +++ b/src/data/nat/order/lemmas.lean @@ -12,7 +12,6 @@ import algebra.group_with_zero.divisibility # Further lemmas about the natural numbers > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/927 > Any changes to this file require a corresponding PR to mathlib4. The distinction between this file and `data.nat.order.basic` is not particularly clear. diff --git a/src/data/nat/psub.lean b/src/data/nat/psub.lean index 1ade5117f4167..325e5b6f66deb 100644 --- a/src/data/nat/psub.lean +++ b/src/data/nat/psub.lean @@ -9,7 +9,6 @@ import data.nat.basic # Partial predecessor and partial subtraction on the natural numbers > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/806 > Any changes to this file require a corresponding PR to mathlib4. The usual definition of natural number subtraction (`nat.sub`) returns 0 as a "garbage value" for diff --git a/src/data/nat/set.lean b/src/data/nat/set.lean index 7e829522e8e34..eab9bf116b1ce 100644 --- a/src/data/nat/set.lean +++ b/src/data/nat/set.lean @@ -10,7 +10,6 @@ import data.set.image ### Recursion on the natural numbers and `set.range` > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/961 > Any changes to this file require a corresponding PR to mathlib4. -/ diff --git a/src/data/nat/units.lean b/src/data/nat/units.lean index 5924e596e8356..5b2818a4c7816 100644 --- a/src/data/nat/units.lean +++ b/src/data/nat/units.lean @@ -6,10 +6,9 @@ Authors: Floris van Doorn, Leonardo de Moura, Jeremy Avigad, Mario Carneiro import data.nat.basic import algebra.group.units -/-! # The units of the natural numbers as a `monoid` and `add_monoid` +/-! # The units of the natural numbers as a `monoid` and `add_monoid` > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/805 > Any changes to this file require a corresponding PR to mathlib4.-/ namespace nat diff --git a/src/data/nat/upto.lean b/src/data/nat/upto.lean index 5a4322ecf74dd..b25775774a18e 100644 --- a/src/data/nat/upto.lean +++ b/src/data/nat/upto.lean @@ -9,7 +9,6 @@ import data.nat.order.basic # `nat.upto` > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/1020 > Any changes to this file require a corresponding PR to mathlib4. `nat.upto p`, with `p` a predicate on `ℕ`, is a subtype of elements `n : ℕ` such that no value diff --git a/src/data/num/basic.lean b/src/data/num/basic.lean index 3a95ab7f0b8e9..91816b9d58233 100644 --- a/src/data/num/basic.lean +++ b/src/data/num/basic.lean @@ -8,7 +8,6 @@ Authors: Leonardo de Moura, Mario Carneiro # Binary representation of integers using inductive types > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/607 > Any changes to this file require a corresponding PR to mathlib4. Note: Unlike in Coq, where this representation is preferred because of diff --git a/src/data/opposite.lean b/src/data/opposite.lean index 9a6beca88e6de..f74f353803b89 100644 --- a/src/data/opposite.lean +++ b/src/data/opposite.lean @@ -9,7 +9,6 @@ import logic.equiv.defs # Opposites > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/650 > Any changes to this file require a corresponding PR to mathlib4. In this file we define a type synonym `opposite α := α`, denoted by `αᵒᵖ` and two synonyms for the diff --git a/src/data/option/basic.lean b/src/data/option/basic.lean index 2a3b9fbd78561..84e4edeb92f0f 100644 --- a/src/data/option/basic.lean +++ b/src/data/option/basic.lean @@ -11,7 +11,6 @@ import tactic.basic # Option of a type > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/493 > Any changes to this file require a corresponding PR to mathlib4. This file develops the basic theory of option types. diff --git a/src/data/option/defs.lean b/src/data/option/defs.lean index ecc410e921e6e..535df54ee80b6 100644 --- a/src/data/option/defs.lean +++ b/src/data/option/defs.lean @@ -8,7 +8,6 @@ Authors: Mario Carneiro # Extra definitions on `option` > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/504 > Any changes to this file require a corresponding PR to mathlib4. This file defines more operations involving `option α`. Lemmas about them are located in other diff --git a/src/data/option/n_ary.lean b/src/data/option/n_ary.lean index 041a0f0df5a8d..007541c0723ca 100644 --- a/src/data/option/n_ary.lean +++ b/src/data/option/n_ary.lean @@ -9,7 +9,6 @@ import data.option.basic # Binary map of options > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/656 > Any changes to this file require a corresponding PR to mathlib4. This file defines the binary map of `option`. This is mostly useful to define pointwise operations diff --git a/src/data/pequiv.lean b/src/data/pequiv.lean index e1451e5f8c07e..1d2c3b745b9e4 100644 --- a/src/data/pequiv.lean +++ b/src/data/pequiv.lean @@ -10,7 +10,6 @@ import data.set.basic # Partial Equivalences > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/1008 > Any changes to this file require a corresponding PR to mathlib4. In this file, we define partial equivalences `pequiv`, which are a bijection between a subset of `α` diff --git a/src/data/pi/algebra.lean b/src/data/pi/algebra.lean index dd41405df5a0e..e3b8a2e710add 100644 --- a/src/data/pi/algebra.lean +++ b/src/data/pi/algebra.lean @@ -16,7 +16,6 @@ import data.prod.basic # Instances and theorems on pi types > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/564 > Any changes to this file require a corresponding PR to mathlib4. This file provides basic definitions and notation instances for Pi types. diff --git a/src/data/pnat/defs.lean b/src/data/pnat/defs.lean index 8c520024fe70f..f08cf158710d0 100644 --- a/src/data/pnat/defs.lean +++ b/src/data/pnat/defs.lean @@ -10,7 +10,6 @@ import algebra.ne_zero # The positive natural numbers > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/604 > Any changes to this file require a corresponding PR to mathlib4. This file contains the definitions, and basic results. diff --git a/src/data/prod/basic.lean b/src/data/prod/basic.lean index 023f589c51477..ca46ef41721bf 100644 --- a/src/data/prod/basic.lean +++ b/src/data/prod/basic.lean @@ -10,7 +10,6 @@ import logic.function.basic # Extra facts about `prod` > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/545 > Any changes to this file require a corresponding PR to mathlib4. This file defines `prod.swap : α × β → β × α` and proves various simple lemmas about `prod`. diff --git a/src/data/prod/lex.lean b/src/data/prod/lex.lean index a0bfe98ca03ff..af22f85079807 100644 --- a/src/data/prod/lex.lean +++ b/src/data/prod/lex.lean @@ -9,7 +9,6 @@ import order.bounded_order # Lexicographic order > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/783 > Any changes to this file require a corresponding PR to mathlib4. This file defines the lexicographic relation for pairs of orders, partial orders and linear orders. diff --git a/src/data/prod/pprod.lean b/src/data/prod/pprod.lean index 14e086446e510..5b1e7f7bfb09f 100644 --- a/src/data/prod/pprod.lean +++ b/src/data/prod/pprod.lean @@ -9,7 +9,6 @@ import logic.basic # Extra facts about `pprod` > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/496 > Any changes to this file require a corresponding PR to mathlib4. -/ diff --git a/src/data/psigma/order.lean b/src/data/psigma/order.lean index 30bdb78fcd518..d148bc589dbbc 100644 --- a/src/data/psigma/order.lean +++ b/src/data/psigma/order.lean @@ -10,7 +10,6 @@ import order.bounded_order # Lexicographic order on a sigma type > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/815 > Any changes to this file require a corresponding PR to mathlib4. This file defines the lexicographic order on `Σₗ' i, α i`. `a` is less than `b` if its summand is diff --git a/src/data/quot.lean b/src/data/quot.lean index b2132e6cbeb0c..eb8bc111569c6 100644 --- a/src/data/quot.lean +++ b/src/data/quot.lean @@ -9,7 +9,6 @@ import logic.relator # Quotient types > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/559 > Any changes to this file require a corresponding PR to mathlib4. This module extends the core library's treatment of quotient types (`init.data.quot`). diff --git a/src/data/rat/init.lean b/src/data/rat/init.lean index e761a4e17465b..988264f5e30f5 100644 --- a/src/data/rat/init.lean +++ b/src/data/rat/init.lean @@ -10,7 +10,6 @@ import logic.basic # The definition of the Rational Numbers > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/575 > Any changes to this file require a corresponding PR to mathlib4. ## Summary diff --git a/src/data/set/basic.lean b/src/data/set/basic.lean index 3f94e8ee93636..b31c965bf21e0 100644 --- a/src/data/set/basic.lean +++ b/src/data/set/basic.lean @@ -10,7 +10,6 @@ import logic.function.iterate # Basic properties of sets > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/892 > Any changes to this file require a corresponding PR to mathlib4. Sets in Lean are homogeneous; all their elements have the same type. Sets whose elements diff --git a/src/data/set/bool_indicator.lean b/src/data/set/bool_indicator.lean index 3b95d84c1a935..3a7718dc53a33 100644 --- a/src/data/set/bool_indicator.lean +++ b/src/data/set/bool_indicator.lean @@ -10,7 +10,6 @@ import data.set.image # Indicator function valued in bool > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/988 > Any changes to this file require a corresponding PR to mathlib4. See also `set.indicator` and `set.piecewise`. diff --git a/src/data/set/image.lean b/src/data/set/image.lean index f1606759efb09..be5dcf78c6a91 100644 --- a/src/data/set/image.lean +++ b/src/data/set/image.lean @@ -9,7 +9,6 @@ import data.set.basic # Images and preimages of sets > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/949 > Any changes to this file require a corresponding PR to mathlib4. ## Main definitions diff --git a/src/data/set/n_ary.lean b/src/data/set/n_ary.lean index 5c43a8082afff..7ebc1528ff099 100644 --- a/src/data/set/n_ary.lean +++ b/src/data/set/n_ary.lean @@ -9,7 +9,6 @@ import data.set.prod # N-ary images of sets > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/969 > Any changes to this file require a corresponding PR to mathlib4. This file defines `finset.image₂`, the binary image of finsets. This is the finset version of diff --git a/src/data/set/opposite.lean b/src/data/set/opposite.lean index 4f5f97cd4f488..9a3786f567a47 100644 --- a/src/data/set/opposite.lean +++ b/src/data/set/opposite.lean @@ -10,7 +10,6 @@ import data.set.image # The opposite of a set > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/983 > Any changes to this file require a corresponding PR to mathlib4. The opposite of a set `s` is simply the set obtained by taking the opposite of each member of `s`. diff --git a/src/data/set/prod.lean b/src/data/set/prod.lean index c8f3bc183cefb..1a5ec550b1712 100644 --- a/src/data/set/prod.lean +++ b/src/data/set/prod.lean @@ -9,7 +9,6 @@ import data.set.image # Sets in product and pi types > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/1004 > Any changes to this file require a corresponding PR to mathlib4. This file defines the product of sets in `α × β` and in `Π i, α i` along with the diagonal of a diff --git a/src/data/set/sigma.lean b/src/data/set/sigma.lean index 24053c92aa2f2..76edf176d89de 100644 --- a/src/data/set/sigma.lean +++ b/src/data/set/sigma.lean @@ -9,7 +9,6 @@ import data.set.image # Sets in sigma types > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/982 > Any changes to this file require a corresponding PR to mathlib4. This file defines `set.sigma`, the indexed sum of sets. diff --git a/src/data/set_like/basic.lean b/src/data/set_like/basic.lean index c89d5dfc8733e..a86e34ae681cf 100644 --- a/src/data/set_like/basic.lean +++ b/src/data/set_like/basic.lean @@ -10,7 +10,6 @@ import tactic.monotonicity.basic # Typeclass for types with a set-like extensionality property > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/993 > Any changes to this file require a corresponding PR to mathlib4. The `has_mem` typeclass is used to let terms of a type have elements. diff --git a/src/data/sigma/basic.lean b/src/data/sigma/basic.lean index 812241abafcb0..33637892d9fc4 100644 --- a/src/data/sigma/basic.lean +++ b/src/data/sigma/basic.lean @@ -12,7 +12,6 @@ import logic.function.basic # Sigma types > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/479 > Any changes to this file require a corresponding PR to mathlib4. This file proves basic results about sigma types. diff --git a/src/data/sigma/lex.lean b/src/data/sigma/lex.lean index 9029366b44c15..c6863bd0ad2c2 100644 --- a/src/data/sigma/lex.lean +++ b/src/data/sigma/lex.lean @@ -10,7 +10,6 @@ import order.rel_classes # Lexicographic order on a sigma type > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/646 > Any changes to this file require a corresponding PR to mathlib4. This defines the lexicographical order of two arbitrary relations on a sigma type and proves some diff --git a/src/data/sigma/order.lean b/src/data/sigma/order.lean index 163d86007be8f..c30c4b0dac88e 100644 --- a/src/data/sigma/order.lean +++ b/src/data/sigma/order.lean @@ -10,7 +10,6 @@ import order.bounded_order # Orders on a sigma type > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/887 > Any changes to this file require a corresponding PR to mathlib4. This file defines two orders on a sigma type: diff --git a/src/data/stream/defs.lean b/src/data/stream/defs.lean index d5c64bf3574fb..c106a14be6e23 100644 --- a/src/data/stream/defs.lean +++ b/src/data/stream/defs.lean @@ -8,7 +8,6 @@ Authors: Leonardo de Moura # Definition of `stream` and functions on streams > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/665 > Any changes to this file require a corresponding PR to mathlib4. A stream `stream α` is an infinite sequence of elements of `α`. One can also think about it as an diff --git a/src/data/subtype.lean b/src/data/subtype.lean index ac204394899cf..5dd326422b4cb 100644 --- a/src/data/subtype.lean +++ b/src/data/subtype.lean @@ -11,7 +11,6 @@ import tactic.simps # Subtypes > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/546 > Any changes to this file require a corresponding PR to mathlib4. This file provides basic API for subtypes, which are defined in core. diff --git a/src/data/sum/basic.lean b/src/data/sum/basic.lean index 3c56efdc03f50..a935d2eec34f2 100644 --- a/src/data/sum/basic.lean +++ b/src/data/sum/basic.lean @@ -10,7 +10,6 @@ import tactic.basic # Disjoint union of types > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/497 > Any changes to this file require a corresponding PR to mathlib4. This file proves basic results about the sum type `α ⊕ β`. diff --git a/src/data/sum/order.lean b/src/data/sum/order.lean index 442b722a519be..c0c1601566d8f 100644 --- a/src/data/sum/order.lean +++ b/src/data/sum/order.lean @@ -9,7 +9,6 @@ import order.hom.basic # Orders on a sum type > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/880 > Any changes to this file require a corresponding PR to mathlib4. This file defines the disjoint sum and the linear (aka lexicographic) sum of two orders and provides diff --git a/src/data/two_pointing.lean b/src/data/two_pointing.lean index 012ae3c0dfb3e..9252b2247314b 100644 --- a/src/data/two_pointing.lean +++ b/src/data/two_pointing.lean @@ -10,7 +10,6 @@ import logic.nontrivial # Two-pointings > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/984 > Any changes to this file require a corresponding PR to mathlib4. This file defines `two_pointing α`, the type of two pointings of `α`. A two-pointing is the data of diff --git a/src/data/ulift.lean b/src/data/ulift.lean index 1c45f97e97983..5384eb374deed 100644 --- a/src/data/ulift.lean +++ b/src/data/ulift.lean @@ -9,7 +9,6 @@ import logic.equiv.basic # Extra lemmas about `ulift` and `plift` > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/703 > Any changes to this file require a corresponding PR to mathlib4. In this file we provide `subsingleton`, `unique`, `decidable_eq`, and `is_empty` instances for diff --git a/src/group_theory/eckmann_hilton.lean b/src/group_theory/eckmann_hilton.lean index 7a25b1ec7ffdf..96230eae00de6 100644 --- a/src/group_theory/eckmann_hilton.lean +++ b/src/group_theory/eckmann_hilton.lean @@ -9,7 +9,6 @@ import algebra.group.defs # Eckmann-Hilton argument > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/626 > Any changes to this file require a corresponding PR to mathlib4. The Eckmann-Hilton argument says that if a type carries two monoid structures that distribute diff --git a/src/group_theory/group_action/defs.lean b/src/group_theory/group_action/defs.lean index fbee61ba9ce03..00aa045ea0f59 100644 --- a/src/group_theory/group_action/defs.lean +++ b/src/group_theory/group_action/defs.lean @@ -13,7 +13,6 @@ import logic.embedding.basic # Definitions of group actions > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/854 > Any changes to this file require a corresponding PR to mathlib4. This file defines a hierarchy of group action type-classes on top of the previously defined diff --git a/src/group_theory/group_action/option.lean b/src/group_theory/group_action/option.lean index 4a08ad259a3d3..9afd75f5ee9cc 100644 --- a/src/group_theory/group_action/option.lean +++ b/src/group_theory/group_action/option.lean @@ -9,7 +9,6 @@ import group_theory.group_action.defs # Option instances for additive and multiplicative actions > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/884 > Any changes to this file require a corresponding PR to mathlib4. This file defines instances for additive and multiplicative actions on `option` type. Scalar diff --git a/src/group_theory/group_action/sigma.lean b/src/group_theory/group_action/sigma.lean index 9648b4e1bc461..fb5378455edc1 100644 --- a/src/group_theory/group_action/sigma.lean +++ b/src/group_theory/group_action/sigma.lean @@ -9,7 +9,6 @@ import group_theory.group_action.defs # Sigma instances for additive and multiplicative actions > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/885 > Any changes to this file require a corresponding PR to mathlib4. This file defines instances for arbitrary sum of additive and multiplicative actions. diff --git a/src/group_theory/group_action/sum.lean b/src/group_theory/group_action/sum.lean index 78f8603467416..253eb5990aa13 100644 --- a/src/group_theory/group_action/sum.lean +++ b/src/group_theory/group_action/sum.lean @@ -9,7 +9,6 @@ import group_theory.group_action.defs # Sum instances for additive and multiplicative actions > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/882 > Any changes to this file require a corresponding PR to mathlib4. This file defines instances for additive and multiplicative actions on the binary `sum` type. diff --git a/src/group_theory/group_action/units.lean b/src/group_theory/group_action/units.lean index 6a9ab262e796d..5aa1f4eaa86f7 100644 --- a/src/group_theory/group_action/units.lean +++ b/src/group_theory/group_action/units.lean @@ -8,7 +8,6 @@ import group_theory.group_action.defs /-! # Group actions on and by `Mˣ` > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/878 > Any changes to this file require a corresponding PR to mathlib4. This file provides the action of a unit on a type `α`, `has_smul Mˣ α`, in the presence of diff --git a/src/logic/basic.lean b/src/logic/basic.lean index 1501e9822e37c..bd88f41f2a834 100644 --- a/src/logic/basic.lean +++ b/src/logic/basic.lean @@ -10,7 +10,6 @@ import tactic.reserved_notation # Basic logic properties > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/484 > Any changes to this file require a corresponding PR to mathlib4. This file is one of the earliest imports in mathlib. diff --git a/src/logic/embedding/basic.lean b/src/logic/embedding/basic.lean index 8334cd6e317a4..06cf19825d790 100644 --- a/src/logic/embedding/basic.lean +++ b/src/logic/embedding/basic.lean @@ -14,7 +14,6 @@ import logic.equiv.basic # Injective functions > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/700 > Any changes to this file require a corresponding PR to mathlib4. -/ diff --git a/src/logic/embedding/set.lean b/src/logic/embedding/set.lean index d219e04aff1b2..60952eac1fe06 100644 --- a/src/logic/embedding/set.lean +++ b/src/logic/embedding/set.lean @@ -10,7 +10,6 @@ import data.set.image # Interactions between embeddings and sets. > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/986 > Any changes to this file require a corresponding PR to mathlib4. -/ diff --git a/src/logic/equiv/basic.lean b/src/logic/equiv/basic.lean index 9fe1b12368638..6628c06c978ca 100644 --- a/src/logic/equiv/basic.lean +++ b/src/logic/equiv/basic.lean @@ -15,7 +15,6 @@ import logic.function.conjugate # Equivalence between types > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/631 > Any changes to this file require a corresponding PR to mathlib4. In this file we continue the work on equivalences begun in `logic/equiv/defs.lean`, defining diff --git a/src/logic/equiv/defs.lean b/src/logic/equiv/defs.lean index d473f09861fdc..abb46c6753b6c 100644 --- a/src/logic/equiv/defs.lean +++ b/src/logic/equiv/defs.lean @@ -10,7 +10,6 @@ import logic.unique # Equivalence between types > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/550 > Any changes to this file require a corresponding PR to mathlib4. In this file we define two types: diff --git a/src/logic/equiv/embedding.lean b/src/logic/equiv/embedding.lean index 547a3e582fb34..ef8d3c64aa9cb 100644 --- a/src/logic/equiv/embedding.lean +++ b/src/logic/equiv/embedding.lean @@ -9,7 +9,6 @@ import logic.embedding.set # Equivalences on embeddings > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/1021 > Any changes to this file require a corresponding PR to mathlib4. This file shows some advanced equivalences on embeddings, useful for constructing larger diff --git a/src/logic/equiv/option.lean b/src/logic/equiv/option.lean index 3988f763e5b46..9c5fe8dc68bde 100644 --- a/src/logic/equiv/option.lean +++ b/src/logic/equiv/option.lean @@ -12,7 +12,6 @@ import logic.equiv.defs # Equivalences for `option α` > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/674 > Any changes to this file require a corresponding PR to mathlib4. diff --git a/src/logic/function/basic.lean b/src/logic/function/basic.lean index fe2b2e20aaffd..f3691cdc4c00e 100644 --- a/src/logic/function/basic.lean +++ b/src/logic/function/basic.lean @@ -11,7 +11,6 @@ import tactic.cache # Miscellaneous function constructions and lemmas > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/511 > Any changes to this file require a corresponding PR to mathlib4. -/ diff --git a/src/logic/function/conjugate.lean b/src/logic/function/conjugate.lean index 8275b8fbb1fd2..85946719b7a8d 100644 --- a/src/logic/function/conjugate.lean +++ b/src/logic/function/conjugate.lean @@ -9,7 +9,6 @@ import logic.function.basic # Semiconjugate and commuting maps > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/533 > Any changes to this file require a corresponding PR to mathlib4. We define the following predicates: diff --git a/src/logic/function/iterate.lean b/src/logic/function/iterate.lean index 92a8484d36892..1f383d7e9bde4 100644 --- a/src/logic/function/iterate.lean +++ b/src/logic/function/iterate.lean @@ -9,7 +9,6 @@ import logic.function.conjugate # Iterations of a function > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/585 > Any changes to this file require a corresponding PR to mathlib4. In this file we prove simple properties of `nat.iterate f n` a.k.a. `f^[n]`: diff --git a/src/logic/is_empty.lean b/src/logic/is_empty.lean index 9af86a2035601..349093c9bcf8b 100644 --- a/src/logic/is_empty.lean +++ b/src/logic/is_empty.lean @@ -10,7 +10,6 @@ import tactic.protected # Types that are empty > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/486 > Any changes to this file require a corresponding PR to mathlib4. In this file we define a typeclass `is_empty`, which expresses that a type has no elements. diff --git a/src/logic/lemmas.lean b/src/logic/lemmas.lean index f868c807a62e3..930d6d764eebc 100644 --- a/src/logic/lemmas.lean +++ b/src/logic/lemmas.lean @@ -13,7 +13,6 @@ import logic.basic # More basic logic properties > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/537 > Any changes to this file require a corresponding PR to mathlib4. A few more logic lemmas. These are in their own file, rather than `logic.basic`, because it is diff --git a/src/logic/nonempty.lean b/src/logic/nonempty.lean index 85efbcdae858b..cb5cca5fcad6f 100644 --- a/src/logic/nonempty.lean +++ b/src/logic/nonempty.lean @@ -9,7 +9,6 @@ import logic.basic # Nonempty types > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/487 > Any changes to this file require a corresponding PR to mathlib4. This file proves a few extra facts about `nonempty`, which is defined in core Lean. diff --git a/src/logic/nontrivial.lean b/src/logic/nontrivial.lean index c8ad9369628b3..99388b9a6b5a1 100644 --- a/src/logic/nontrivial.lean +++ b/src/logic/nontrivial.lean @@ -12,7 +12,6 @@ import logic.unique # Nontrivial types > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/547 > Any changes to this file require a corresponding PR to mathlib4. A type is *nontrivial* if it contains at least two elements. This is useful in particular for rings diff --git a/src/logic/pairwise.lean b/src/logic/pairwise.lean index f7d9b833b3276..4835ef323ddfc 100644 --- a/src/logic/pairwise.lean +++ b/src/logic/pairwise.lean @@ -10,7 +10,6 @@ import tactic.basic # Relations holding pairwise > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/622 > Any changes to this file require a corresponding PR to mathlib4. This file defines pairwise relations. diff --git a/src/logic/relation.lean b/src/logic/relation.lean index a30950cd28fd3..8467caa176c36 100644 --- a/src/logic/relation.lean +++ b/src/logic/relation.lean @@ -10,7 +10,6 @@ import logic.relator # Relation closures > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/565 > Any changes to this file require a corresponding PR to mathlib4. This file defines the reflexive, transitive, and reflexive transitive closures of relations. diff --git a/src/logic/relator.lean b/src/logic/relator.lean index 013d5ded9cf00..e2d4251b193b5 100644 --- a/src/logic/relator.lean +++ b/src/logic/relator.lean @@ -10,7 +10,6 @@ import logic.basic # Relator for functions, pairs, sums, and lists. > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/385 > Any changes to this file require a corresponding PR to mathlib4. -/ diff --git a/src/logic/unique.lean b/src/logic/unique.lean index c77b82822acaa..5384a49f64e38 100644 --- a/src/logic/unique.lean +++ b/src/logic/unique.lean @@ -10,7 +10,6 @@ import logic.is_empty # Types with a unique term > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/512 > Any changes to this file require a corresponding PR to mathlib4. In this file we define a typeclass `unique`, diff --git a/src/order/antisymmetrization.lean b/src/order/antisymmetrization.lean index 78b4db2f22500..33d319b56e7d7 100644 --- a/src/order/antisymmetrization.lean +++ b/src/order/antisymmetrization.lean @@ -10,7 +10,6 @@ import logic.relation # Turning a preorder into a partial order > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/876 > Any changes to this file require a corresponding PR to mathlib4. This file allows to make a preorder into a partial order by quotienting out the elements `a`, `b` diff --git a/src/order/basic.lean b/src/order/basic.lean index 859289add65e5..7d174e19e8aff 100644 --- a/src/order/basic.lean +++ b/src/order/basic.lean @@ -10,7 +10,6 @@ import data.subtype # Basic definitions about `≤` and `<` > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/556 > Any changes to this file require a corresponding PR to mathlib4. This file proves basic results about orders, provides extensive dot notation, defines useful order diff --git a/src/order/boolean_algebra.lean b/src/order/boolean_algebra.lean index 6bb520e1d9aaf..f0c33b28e318f 100644 --- a/src/order/boolean_algebra.lean +++ b/src/order/boolean_algebra.lean @@ -9,7 +9,6 @@ import order.heyting.basic # (Generalized) Boolean algebras > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/794 > Any changes to this file require a corresponding PR to mathlib4. A Boolean algebra is a bounded distributive lattice with a complement operator. Boolean algebras diff --git a/src/order/bounded_order.lean b/src/order/bounded_order.lean index c04634f8e9f6f..c9c998236dcf0 100644 --- a/src/order/bounded_order.lean +++ b/src/order/bounded_order.lean @@ -10,7 +10,6 @@ import data.option.basic # ⊤ and ⊥, bounded lattices and variants > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/697 > Any changes to this file require a corresponding PR to mathlib4. This file defines top and bottom elements (greatest and least elements) of a type, the bounded diff --git a/src/order/compare.lean b/src/order/compare.lean index edd8b1beed6de..92fb96f856652 100644 --- a/src/order/compare.lean +++ b/src/order/compare.lean @@ -9,7 +9,6 @@ import order.synonym # Comparison > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/569 > Any changes to this file require a corresponding PR to mathlib4. This file provides basic results about orderings and comparison in linear orders. diff --git a/src/order/directed.lean b/src/order/directed.lean index b52108f4f7b41..1d84b95013f1f 100644 --- a/src/order/directed.lean +++ b/src/order/directed.lean @@ -11,7 +11,6 @@ import order.max # Directed indexed families and sets > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/963 > Any changes to this file require a corresponding PR to mathlib4. This file defines directed indexed families and directed sets. An indexed family/set is diff --git a/src/order/disjoint.lean b/src/order/disjoint.lean index d9e3f8daf84c9..bdff98cb1479d 100644 --- a/src/order/disjoint.lean +++ b/src/order/disjoint.lean @@ -9,7 +9,6 @@ import order.bounded_order # Disjointness and complements > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/773 > Any changes to this file require a corresponding PR to mathlib4. This file defines `disjoint`, `codisjoint`, and the `is_compl` predicate. diff --git a/src/order/game_add.lean b/src/order/game_add.lean index dc17266a04583..f7207da171cae 100644 --- a/src/order/game_add.lean +++ b/src/order/game_add.lean @@ -10,7 +10,6 @@ import logic.relation # Game addition relation > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/645 > Any changes to this file require a corresponding PR to mathlib4. This file defines, given relations `rα : α → α → Prop` and `rβ : β → β → Prop`, a relation diff --git a/src/order/heyting/basic.lean b/src/order/heyting/basic.lean index 8d400e2e71562..8d190745d9817 100644 --- a/src/order/heyting/basic.lean +++ b/src/order/heyting/basic.lean @@ -9,7 +9,6 @@ import order.prop_instances # Heyting algebras > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/793 > Any changes to this file require a corresponding PR to mathlib4. This file defines Heyting, co-Heyting and bi-Heyting algebras. diff --git a/src/order/heyting/boundary.lean b/src/order/heyting/boundary.lean index 23c95debed7c9..e267bffe81795 100644 --- a/src/order/heyting/boundary.lean +++ b/src/order/heyting/boundary.lean @@ -9,7 +9,6 @@ import order.boolean_algebra # Co-Heyting boundary > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/844 > Any changes to this file require a corresponding PR to mathlib4. The boundary of an element of a co-Heyting algebra is the intersection of its Heyting negation with diff --git a/src/order/hom/basic.lean b/src/order/hom/basic.lean index 9ba2f6867f416..ee8351a2f7e71 100644 --- a/src/order/hom/basic.lean +++ b/src/order/hom/basic.lean @@ -13,7 +13,6 @@ import order.disjoint # Order homomorphisms > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/804 > Any changes to this file require a corresponding PR to mathlib4. This file defines order homomorphisms, which are bundled monotone functions. A preorder diff --git a/src/order/iterate.lean b/src/order/iterate.lean index 70921d0907b83..1dfdb3b58f77d 100644 --- a/src/order/iterate.lean +++ b/src/order/iterate.lean @@ -10,7 +10,6 @@ import order.monotone.basic # Inequalities on iterates > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/648 > Any changes to this file require a corresponding PR to mathlib4. In this file we prove some inequalities comparing `f^[n] x` and `g^[n] x` where `f` and `g` are diff --git a/src/order/lattice.lean b/src/order/lattice.lean index 7961e75b4779d..e41a99100a96b 100644 --- a/src/order/lattice.lean +++ b/src/order/lattice.lean @@ -11,7 +11,6 @@ import tactic.pi_instances # (Semi-)lattices > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/642 > Any changes to this file require a corresponding PR to mathlib4. Semilattices are partially ordered sets with join (greatest lower bound, or `sup`) or diff --git a/src/order/max.lean b/src/order/max.lean index 54860a98e8e62..ad59fa0fa0b6f 100644 --- a/src/order/max.lean +++ b/src/order/max.lean @@ -9,7 +9,6 @@ import order.synonym # Minimal/maximal and bottom/top elements > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/567 > Any changes to this file require a corresponding PR to mathlib4. This file defines predicates for elements to be minimal/maximal or bottom/top and typeclasses diff --git a/src/order/min_max.lean b/src/order/min_max.lean index b8652ef332655..6703101a03467 100644 --- a/src/order/min_max.lean +++ b/src/order/min_max.lean @@ -9,7 +9,6 @@ import order.lattice # `max` and `min` > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/728 > Any changes to this file require a corresponding PR to mathlib4. This file proves basic properties about maxima and minima on a `linear_order`. diff --git a/src/order/monotone/basic.lean b/src/order/monotone/basic.lean index 7f1a05ce9834f..e107e23328433 100644 --- a/src/order/monotone/basic.lean +++ b/src/order/monotone/basic.lean @@ -11,7 +11,6 @@ import order.rel_classes # Monotonicity > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/591 > Any changes to this file require a corresponding PR to mathlib4. This file defines (strictly) monotone/antitone functions. Contrary to standard mathematical usage, diff --git a/src/order/prop_instances.lean b/src/order/prop_instances.lean index b9269d8bc8337..e27289e1d8153 100644 --- a/src/order/prop_instances.lean +++ b/src/order/prop_instances.lean @@ -11,7 +11,6 @@ import order.with_bot # The order on `Prop` > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/792 > Any changes to this file require a corresponding PR to mathlib4. Instances on `Prop` such as `distrib_lattice`, `bounded_order`, `linear_order`. diff --git a/src/order/rel_classes.lean b/src/order/rel_classes.lean index 44a54cd1d1586..4ebbf966321d7 100644 --- a/src/order/rel_classes.lean +++ b/src/order/rel_classes.lean @@ -10,7 +10,6 @@ import logic.is_empty # Unbundled relation classes > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/560 > Any changes to this file require a corresponding PR to mathlib4. In this file we prove some properties of `is_*` classes defined in `init.algebra.classes`. The main diff --git a/src/order/rel_iso/basic.lean b/src/order/rel_iso/basic.lean index 5975e5bc02c64..f2bb56d779227 100644 --- a/src/order/rel_iso/basic.lean +++ b/src/order/rel_iso/basic.lean @@ -11,7 +11,6 @@ import order.rel_classes # Relation homomorphisms, embeddings, isomorphisms > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/772 > Any changes to this file require a corresponding PR to mathlib4. This file defines relation homomorphisms, embeddings, isomorphisms and order embeddings and diff --git a/src/order/rel_iso/group.lean b/src/order/rel_iso/group.lean index 5a1851128a2c5..0cc4b0c71b194 100644 --- a/src/order/rel_iso/group.lean +++ b/src/order/rel_iso/group.lean @@ -10,7 +10,6 @@ import order.rel_iso.basic # Relation isomorphisms form a group > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/813 > Any changes to this file require a corresponding PR to mathlib4. -/ diff --git a/src/order/rel_iso/set.lean b/src/order/rel_iso/set.lean index 91a8430539550..86649ec7ef5a0 100644 --- a/src/order/rel_iso/set.lean +++ b/src/order/rel_iso/set.lean @@ -10,7 +10,6 @@ import logic.embedding.set # Interactions between relation homomorphisms and sets > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/1005 > Any changes to this file require a corresponding PR to mathlib4. It is likely that there are better homes for many of these statement, diff --git a/src/order/symm_diff.lean b/src/order/symm_diff.lean index 803677f2be73b..03b12b22de22a 100644 --- a/src/order/symm_diff.lean +++ b/src/order/symm_diff.lean @@ -11,7 +11,6 @@ import logic.equiv.basic # Symmetric difference and bi-implication > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/842 > Any changes to this file require a corresponding PR to mathlib4. This file defines the symmetric difference and bi-implication operators in (co-)Heyting algebras. diff --git a/src/order/synonym.lean b/src/order/synonym.lean index 1de69a1fcc315..f04b5957db6d0 100644 --- a/src/order/synonym.lean +++ b/src/order/synonym.lean @@ -11,7 +11,6 @@ import order.basic # Type synonyms > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/562 > Any changes to this file require a corresponding PR to mathlib4. This file provides two type synonyms for order theory: diff --git a/src/order/well_founded.lean b/src/order/well_founded.lean index 101742ceaa75f..46f5f29ab47da 100644 --- a/src/order/well_founded.lean +++ b/src/order/well_founded.lean @@ -10,7 +10,6 @@ import data.set.image # Well-founded relations > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/970 > Any changes to this file require a corresponding PR to mathlib4. A relation is well-founded if it can be used for induction: for each `x`, `(∀ y, r y x → P y) → P x` diff --git a/src/order/with_bot.lean b/src/order/with_bot.lean index 61431d9664c96..93d3cd2093e2c 100644 --- a/src/order/with_bot.lean +++ b/src/order/with_bot.lean @@ -9,7 +9,6 @@ import order.bounded_order # `with_bot`, `with_top` > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/776 > Any changes to this file require a corresponding PR to mathlib4. Adding a `bot` or a `top` to an order. diff --git a/src/ring_theory/coprime/basic.lean b/src/ring_theory/coprime/basic.lean index 940441dcd4fec..1a3f23e2a122f 100644 --- a/src/ring_theory/coprime/basic.lean +++ b/src/ring_theory/coprime/basic.lean @@ -10,7 +10,6 @@ import group_theory.group_action.units # Coprime elements of a ring > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> https://github.com/leanprover-community/mathlib4/pull/899 > Any changes to this file require a corresponding PR to mathlib4. ## Main definitions From c3442db180528e8662b206e474a22613876f8d56 Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Mon, 2 Jan 2023 14:48:22 +0000 Subject: [PATCH 0127/1291] chore(ring_theory/dedekind_domain/ideal): remove fractional_ideal in open namespaces (#18033) --- src/ring_theory/dedekind_domain/ideal.lean | 142 +++++++++------------ src/ring_theory/fractional_ideal.lean | 35 ++--- 2 files changed, 77 insertions(+), 100 deletions(-) diff --git a/src/ring_theory/dedekind_domain/ideal.lean b/src/ring_theory/dedekind_domain/ideal.lean index 62a116f164c96..6ac888ca97883 100644 --- a/src/ring_theory/dedekind_domain/ideal.lean +++ b/src/ring_theory/dedekind_domain/ideal.lean @@ -66,11 +66,10 @@ noncomputable instance : has_inv (fractional_ideal R₁⁰ K) := ⟨λ I, 1 / I lemma inv_eq : I⁻¹ = 1 / I := rfl -lemma inv_zero' : (0 : fractional_ideal R₁⁰ K)⁻¹ = 0 := fractional_ideal.div_zero +lemma inv_zero' : (0 : fractional_ideal R₁⁰ K)⁻¹ = 0 := div_zero lemma inv_nonzero {J : fractional_ideal R₁⁰ K} (h : J ≠ 0) : -J⁻¹ = ⟨(1 : fractional_ideal R₁⁰ K) / J, fractional_ideal.fractional_div_of_nonzero h⟩ := -fractional_ideal.div_nonzero _ +J⁻¹ = ⟨(1 : fractional_ideal R₁⁰ K) / J, fractional_div_of_nonzero h⟩ := div_nonzero _ lemma coe_inv_of_nonzero {J : fractional_ideal R₁⁰ K} (h : J ≠ 0) : (↑J⁻¹ : submodule R₁ K) = is_localization.coe_submodule K ⊤ / J := @@ -78,64 +77,57 @@ by { rwa inv_nonzero _, refl, assumption } variables {K} -lemma mem_inv_iff (hI : I ≠ 0) {x : K} : - x ∈ I⁻¹ ↔ ∀ y ∈ I, x * y ∈ (1 : fractional_ideal R₁⁰ K) := -fractional_ideal.mem_div_iff_of_nonzero hI +lemma mem_inv_iff (hI : I ≠ 0) {x : K} : x ∈ I⁻¹ ↔ ∀ y ∈ I, x * y ∈ (1 : fractional_ideal R₁⁰ K) := +mem_div_iff_of_nonzero hI -lemma inv_anti_mono (hI : I ≠ 0) (hJ : J ≠ 0) (hIJ : I ≤ J) : - J⁻¹ ≤ I⁻¹ := +lemma inv_anti_mono (hI : I ≠ 0) (hJ : J ≠ 0) (hIJ : I ≤ J) : J⁻¹ ≤ I⁻¹ := λ x, by { simp only [mem_inv_iff hI, mem_inv_iff hJ], exact λ h y hy, h y (hIJ hy) } lemma le_self_mul_inv {I : fractional_ideal R₁⁰ K} (hI : I ≤ (1 : fractional_ideal R₁⁰ K)) : I ≤ I * I⁻¹ := -fractional_ideal.le_self_mul_one_div hI +le_self_mul_one_div hI variables (K) -lemma coe_ideal_le_self_mul_inv (I : ideal R₁) : - (I : fractional_ideal R₁⁰ K) ≤ I * I⁻¹ := -le_self_mul_inv fractional_ideal.coe_ideal_le_one +lemma coe_ideal_le_self_mul_inv (I : ideal R₁) : (I : fractional_ideal R₁⁰ K) ≤ I * I⁻¹ := +le_self_mul_inv coe_ideal_le_one /-- `I⁻¹` is the inverse of `I` if `I` has an inverse. -/ -theorem right_inverse_eq (I J : fractional_ideal R₁⁰ K) (h : I * J = 1) : - J = I⁻¹ := +theorem right_inverse_eq (I J : fractional_ideal R₁⁰ K) (h : I * J = 1) : J = I⁻¹ := begin - have hI : I ≠ 0 := fractional_ideal.ne_zero_of_mul_eq_one I J h, + have hI : I ≠ 0 := ne_zero_of_mul_eq_one I J h, suffices h' : I * (1 / I) = 1, { exact (congr_arg units.inv $ @units.ext _ _ (units.mk_of_mul_eq_one _ _ h) (units.mk_of_mul_eq_one _ _ h') rfl) }, apply le_antisymm, - { apply fractional_ideal.mul_le.mpr _, + { apply mul_le.mpr _, intros x hx y hy, rw mul_comm, - exact (fractional_ideal.mem_div_iff_of_nonzero hI).mp hy x hx }, + exact (mem_div_iff_of_nonzero hI).mp hy x hx }, rw ← h, - apply fractional_ideal.mul_left_mono I, - apply (fractional_ideal.le_div_iff_of_nonzero hI).mpr _, + apply mul_left_mono I, + apply (le_div_iff_of_nonzero hI).mpr _, intros y hy x hx, rw mul_comm, - exact fractional_ideal.mul_mem_mul hx hy + exact mul_mem_mul hx hy end -theorem mul_inv_cancel_iff {I : fractional_ideal R₁⁰ K} : - I * I⁻¹ = 1 ↔ ∃ J, I * J = 1 := +theorem mul_inv_cancel_iff {I : fractional_ideal R₁⁰ K} : I * I⁻¹ = 1 ↔ ∃ J, I * J = 1 := ⟨λ h, ⟨I⁻¹, h⟩, λ ⟨J, hJ⟩, by rwa ← right_inverse_eq K I J hJ⟩ -lemma mul_inv_cancel_iff_is_unit {I : fractional_ideal R₁⁰ K} : - I * I⁻¹ = 1 ↔ is_unit I := +lemma mul_inv_cancel_iff_is_unit {I : fractional_ideal R₁⁰ K} : I * I⁻¹ = 1 ↔ is_unit I := (mul_inv_cancel_iff K).trans is_unit_iff_exists_inv.symm variables {K' : Type*} [field K'] [algebra R₁ K'] [is_fraction_ring R₁ K'] @[simp] lemma map_inv (I : fractional_ideal R₁⁰ K) (h : K ≃ₐ[R₁] K') : (I⁻¹).map (h : K →ₐ[R₁] K') = (I.map h)⁻¹ := -by rw [inv_eq, fractional_ideal.map_div, fractional_ideal.map_one, inv_eq] +by rw [inv_eq, map_div, map_one, inv_eq] open submodule submodule.is_principal -@[simp] lemma span_singleton_inv (x : K) : - (fractional_ideal.span_singleton R₁⁰ x)⁻¹ = fractional_ideal.span_singleton _ (x⁻¹) := -fractional_ideal.one_div_span_singleton x +@[simp] lemma span_singleton_inv (x : K) : (span_singleton R₁⁰ x)⁻¹ = span_singleton _ x⁻¹ := +one_div_span_singleton x lemma coe_ideal_span_singleton_mul_inv {x : R₁} (hx : x ≠ 0) : ((ideal.span {x} : ideal R₁) : fractional_ideal R₁⁰ K) * (ideal.span {x} : ideal R₁)⁻¹ = 1 := @@ -149,24 +141,20 @@ by rw [mul_comm, coe_ideal_span_singleton_mul_inv K hx] lemma mul_generator_self_inv {R₁ : Type*} [comm_ring R₁] [algebra R₁ K] [is_localization R₁⁰ K] (I : fractional_ideal R₁⁰ K) [submodule.is_principal (I : submodule R₁ K)] (h : I ≠ 0) : - I * fractional_ideal.span_singleton _ (generator (I : submodule R₁ K))⁻¹ = 1 := + I * span_singleton _ (generator (I : submodule R₁ K))⁻¹ = 1 := begin -- Rewrite only the `I` that appears alone. - conv_lhs { congr, rw fractional_ideal.eq_span_singleton_of_principal I }, - rw [fractional_ideal.span_singleton_mul_span_singleton, mul_inv_cancel, - fractional_ideal.span_singleton_one], + conv_lhs { congr, rw eq_span_singleton_of_principal I }, + rw [span_singleton_mul_span_singleton, mul_inv_cancel, span_singleton_one], intro generator_I_eq_zero, apply h, - rw [fractional_ideal.eq_span_singleton_of_principal I, generator_I_eq_zero, - fractional_ideal.span_singleton_zero] + rw [eq_span_singleton_of_principal I, generator_I_eq_zero, span_singleton_zero] end lemma invertible_of_principal (I : fractional_ideal R₁⁰ K) - [submodule.is_principal (I : submodule R₁ K)] (h : I ≠ 0) : - I * I⁻¹ = 1 := -(fractional_ideal.mul_div_self_cancel_iff).mpr - ⟨fractional_ideal.span_singleton _ (generator (I : submodule R₁ K))⁻¹, - mul_generator_self_inv _ I h⟩ + [submodule.is_principal (I : submodule R₁ K)] (h : I ≠ 0) : I * I⁻¹ = 1 := +(mul_div_self_cancel_iff).mpr + ⟨span_singleton _ (generator (I : submodule R₁ K))⁻¹, mul_generator_self_inv _ I h⟩ lemma invertible_iff_generator_nonzero (I : fractional_ideal R₁⁰ K) [submodule.is_principal (I : submodule R₁ K)] : @@ -174,15 +162,14 @@ lemma invertible_iff_generator_nonzero (I : fractional_ideal R₁⁰ K) begin split, { intros hI hg, - apply fractional_ideal.ne_zero_of_mul_eq_one _ _ hI, - rw [fractional_ideal.eq_span_singleton_of_principal I, hg, - fractional_ideal.span_singleton_zero] }, + apply ne_zero_of_mul_eq_one _ _ hI, + rw [eq_span_singleton_of_principal I, hg, span_singleton_zero] }, { intro hg, apply invertible_of_principal, - rw [fractional_ideal.eq_span_singleton_of_principal I], + rw [eq_span_singleton_of_principal I], intro hI, - have := fractional_ideal.mem_span_singleton_self _ (generator (I : submodule R₁ K)), - rw [hI, fractional_ideal.mem_zero_iff] at this, + have := mem_span_singleton_self _ (generator (I : submodule R₁ K)), + rw [hI, mem_zero_iff] at this, contradiction } end @@ -190,16 +177,15 @@ lemma is_principal_inv (I : fractional_ideal R₁⁰ K) [submodule.is_principal (I : submodule R₁ K)] (h : I ≠ 0) : submodule.is_principal (I⁻¹).1 := begin - rw [fractional_ideal.val_eq_coe, fractional_ideal.is_principal_iff], + rw [val_eq_coe, is_principal_iff], use (generator (I : submodule R₁ K))⁻¹, - have hI : I * fractional_ideal.span_singleton _ ((generator (I : submodule R₁ K))⁻¹) = 1, + have hI : I * span_singleton _ ((generator (I : submodule R₁ K))⁻¹) = 1, apply mul_generator_self_inv _ I h, - exact (right_inverse_eq _ I (fractional_ideal.span_singleton _ - ((generator (I : submodule R₁ K))⁻¹)) hI).symm + exact (right_inverse_eq _ I (span_singleton _ ((generator (I : submodule R₁ K))⁻¹)) hI).symm end noncomputable instance : inv_one_class (fractional_ideal R₁⁰ K) := -{ inv_one := fractional_ideal.div_one, +{ inv_one := div_one, ..fractional_ideal.has_one, ..fractional_ideal.has_inv K } @@ -221,10 +207,9 @@ open fractional_ideal variables {R A K} lemma is_dedekind_domain_inv_iff [algebra A K] [is_fraction_ring A K] : - is_dedekind_domain_inv A ↔ - (∀ I ≠ (⊥ : fractional_ideal A⁰ K), I * I⁻¹ = 1) := + is_dedekind_domain_inv A ↔ (∀ I ≠ (⊥ : fractional_ideal A⁰ K), I * I⁻¹ = 1) := begin - let h := fractional_ideal.map_equiv (fraction_ring.alg_equiv A K), + let h := map_equiv (fraction_ring.alg_equiv A K), refine h.to_equiv.forall_congr (λ I, _), rw ← h.to_equiv.apply_eq_iff_eq, simp [is_dedekind_domain_inv, show ⇑h.to_equiv = h, from rfl], @@ -236,7 +221,7 @@ lemma fractional_ideal.adjoin_integral_eq_one_of_is_unit [algebra A K] [is_fract begin set I := adjoin_integral A⁰ x hx, have mul_self : I * I = I, - { apply fractional_ideal.coe_to_submodule_injective, simp }, + { apply coe_to_submodule_injective, simp }, convert congr_arg (* I⁻¹) mul_self; simp only [(mul_inv_cancel_iff_is_unit K).mpr hI, mul_assoc, mul_one], end @@ -272,7 +257,7 @@ begin -- `A[x]` (which is a fractional ideal) is in fact equal to `A`. refine ⟨λ x hx, _⟩, rw [← set.mem_range, ← algebra.mem_bot, ← subalgebra.mem_to_submodule, algebra.to_submodule_bot, - ← coe_span_singleton A⁰ (1 : fraction_ring A), fractional_ideal.span_singleton_one, + ← coe_span_singleton A⁰ (1 : fraction_ring A), span_singleton_one, ← fractional_ideal.adjoin_integral_eq_one_of_is_unit x hx (h.is_unit _)], { exact mem_adjoin_integral_self A⁰ x hx }, { exact λ h, one_ne_zero (eq_zero_iff.mp h 1 (subalgebra.one_mem _)) }, @@ -295,7 +280,7 @@ begin -- In particular, we'll show `M⁻¹ * P ≤ P` suffices : (M⁻¹ * P : fractional_ideal A⁰ (fraction_ring A)) ≤ P, - { rw [eq_top_iff, ← coe_ideal_le_coe_ideal (fraction_ring A), fractional_ideal.coe_ideal_top], + { rw [eq_top_iff, ← coe_ideal_le_coe_ideal (fraction_ring A), coe_ideal_top], calc (1 : fractional_ideal A⁰ (fraction_ring A)) = _ * _ * _ : _ ... ≤ _ * _ : mul_right_mono (P⁻¹ * M : fractional_ideal A⁰ (fraction_ring A)) this ... = M : _, @@ -307,13 +292,13 @@ begin intros x hx, have le_one : (M⁻¹ * P : fractional_ideal A⁰ (fraction_ring A)) ≤ 1, { rw [← h.inv_mul_eq_one M'_ne], - exact fractional_ideal.mul_left_mono _ ((coe_ideal_le_coe_ideal (fraction_ring A)).mpr hM.le) }, + exact mul_left_mono _ ((coe_ideal_le_coe_ideal (fraction_ring A)).mpr hM.le) }, obtain ⟨y, hy, rfl⟩ := (mem_coe_ideal _).mp (le_one hx), -- Since `M` is strictly greater than `P`, let `z ∈ M \ P`. obtain ⟨z, hzM, hzp⟩ := set_like.exists_of_lt hM, -- We have `z * y ∈ M * (M⁻¹ * P) = P`. - have zy_mem := fractional_ideal.mul_mem_mul (mem_coe_ideal_of_mem A⁰ hzM) hx, + have zy_mem := mul_mem_mul (mem_coe_ideal_of_mem A⁰ hzM) hx, rw [← ring_hom.map_mul, ← mul_assoc, h.mul_inv_eq_one M'_ne, one_mul] at zy_mem, obtain ⟨zy, hzy, zy_eq⟩ := (mem_coe_ideal A⁰).mp zy_mem, rw is_fraction_ring.injective A (fraction_ring A) zy_eq at hzy, @@ -385,7 +370,7 @@ begin have hM0 := (M.bot_lt_of_maximal hNF).ne', obtain ⟨x, hxM, hx1⟩ := this hM, refine ⟨x, inv_anti_mono _ _ ((coe_ideal_le_coe_ideal _).mpr hIM) hxM, hx1⟩; - apply fractional_ideal.coe_ideal_ne_zero; assumption }, + apply coe_ideal_ne_zero; assumption }, -- Let `a` be a nonzero element of `M` and `J` the ideal generated by `a`. intros M hM, @@ -409,9 +394,9 @@ begin (λ h, hbJ $ h.symm ▸ J.zero_mem), -- Then `b a⁻¹ : K` is in `M⁻¹` but not in `1`. refine ⟨algebra_map A K b * (algebra_map A K a)⁻¹, (mem_inv_iff _).mpr _, _⟩, - { exact (fractional_ideal.coe_to_fractional_ideal_ne_zero le_rfl).mpr hM0.ne' }, + { exact (coe_to_fractional_ideal_ne_zero le_rfl).mpr hM0.ne' }, { rintro y₀ hy₀, - obtain ⟨y, h_Iy, rfl⟩ := (fractional_ideal.mem_coe_ideal _).mp hy₀, + obtain ⟨y, h_Iy, rfl⟩ := (mem_coe_ideal _).mp hy₀, rw [mul_comm, ← mul_assoc, ← ring_hom.map_mul], have h_yb : y * b ∈ J, { apply hle, @@ -420,8 +405,8 @@ begin rw ideal.mem_span_singleton' at h_yb, rcases h_yb with ⟨c, hc⟩, rw [← hc, ring_hom.map_mul, mul_assoc, mul_inv_cancel hnz_fa, mul_one], - apply fractional_ideal.coe_mem_one }, - { refine mt (fractional_ideal.mem_one_iff _).mp _, + apply coe_mem_one }, + { refine mt (mem_one_iff _).mp _, rintros ⟨x', h₂_abs⟩, rw [← div_eq_mul_inv, eq_div_iff_mul_eq hnz_fa, ← ring_hom.map_mul] at h₂_abs, have := ideal.mem_span_singleton'.mpr ⟨x', is_fraction_ring.injective A K h₂_abs⟩, @@ -431,7 +416,7 @@ end lemma one_mem_inv_coe_ideal {I : ideal A} (hI : I ≠ ⊥) : (1 : K) ∈ (I : fractional_ideal A⁰ K)⁻¹ := begin - rw mem_inv_iff (fractional_ideal.coe_ideal_ne_zero hI), + rw mem_inv_iff (coe_ideal_ne_zero hI), intros y hy, rw one_mul, exact coe_ideal_le_one hy, @@ -478,12 +463,12 @@ begin -- We'll show `1 ≤ J⁻¹ = (I * I⁻¹)⁻¹ ≤ 1`. apply mul_inv_cancel_of_le_one hI0, by_cases hJ0 : (I * I⁻¹ : fractional_ideal A⁰ K) = 0, - { rw [hJ0, inv_zero'], exact fractional_ideal.zero_le _ }, + { rw [hJ0, inv_zero'], exact zero_le _ }, intros x hx, -- In particular, we'll show all `x ∈ J⁻¹` are integral. suffices : x ∈ integral_closure A K, { rwa [is_integrally_closed.integral_closure_eq_bot, algebra.mem_bot, set.mem_range, - ← fractional_ideal.mem_one_iff] at this; + ← mem_one_iff] at this; assumption }, -- For that, we'll find a subalgebra that is f.g. as a module and contains `x`. -- `A` is a noetherian ring, so we just need to find a subalgebra between `{x}` and `I⁻¹`. @@ -491,14 +476,14 @@ begin have x_mul_mem : ∀ b ∈ (I⁻¹ : fractional_ideal A⁰ K), x * b ∈ (I⁻¹ : fractional_ideal A⁰ K), { intros b hb, rw mem_inv_iff at ⊢ hx, - swap, { exact fractional_ideal.coe_ideal_ne_zero hI0 }, + swap, { exact coe_ideal_ne_zero hI0 }, swap, { exact hJ0 }, simp only [mul_assoc, mul_comm b] at ⊢ hx, intros y hy, - exact hx _ (fractional_ideal.mul_mem_mul hy hb) }, + exact hx _ (mul_mem_mul hy hb) }, -- It turns out the subalgebra consisting of all `p(x)` for `p : A[X]` works. refine ⟨alg_hom.range (polynomial.aeval x : A[X] →ₐ[A] K), - is_noetherian_submodule.mp (fractional_ideal.is_noetherian I⁻¹) _ (λ y hy, _), + is_noetherian_submodule.mp (is_noetherian I⁻¹) _ (λ y hy, _), ⟨polynomial.X, polynomial.aeval_X x⟩⟩, obtain ⟨p, rfl⟩ := (alg_hom.mem_range _).mp hy, rw polynomial.aeval_eq_sum_range, @@ -526,11 +511,10 @@ begin exact ⟨span_singleton A⁰ (algebra_map _ _ a) * J⁻¹, h₂⟩ }, subst hJ, rw [mul_assoc, mul_left_comm (J : fractional_ideal A⁰ K), coe_ideal_mul_inv, mul_one, - fractional_ideal.span_singleton_mul_span_singleton, inv_mul_cancel, - fractional_ideal.span_singleton_one], + span_singleton_mul_span_singleton, inv_mul_cancel, span_singleton_one], { exact mt ((injective_iff_map_eq_zero (algebra_map A K)).mp (is_fraction_ring.injective A K) _) ha }, - { exact fractional_ideal.coe_ideal_ne_zero_iff.mp (right_ne_zero_of_mul hne) } + { exact coe_ideal_ne_zero_iff.mp (right_ne_zero_of_mul hne) } end lemma mul_right_le_iff [is_dedekind_domain A] {J : fractional_ideal A⁰ K} @@ -545,7 +529,7 @@ end lemma mul_left_le_iff [is_dedekind_domain A] {J : fractional_ideal A⁰ K} (hJ : J ≠ 0) {I I'} : J * I ≤ J * I' ↔ I ≤ I' := -by convert fractional_ideal.mul_right_le_iff hJ using 1; simp only [mul_comm] +by convert mul_right_le_iff hJ using 1; simp only [mul_comm] lemma mul_right_strict_mono [is_dedekind_domain A] {I : fractional_ideal A⁰ K} (hI : I ≠ 0) : strict_mono (* I) := @@ -629,15 +613,15 @@ lemma ideal.dvd_iff_le {I J : ideal A} : (I ∣ J) ↔ J ≤ I := { have hJ : J = ⊥, { rwa [hI, ← eq_bot_iff] at h }, rw [hI, hJ] }, have hI' : (I : fractional_ideal A⁰ (fraction_ring A)) ≠ 0 := - (fractional_ideal.coe_to_fractional_ideal_ne_zero (le_refl (non_zero_divisors A))).mpr hI, + (coe_to_fractional_ideal_ne_zero (le_refl (non_zero_divisors A))).mpr hI, have : (I : fractional_ideal A⁰ (fraction_ring A))⁻¹ * J ≤ 1 := le_trans - (fractional_ideal.mul_left_mono (↑I)⁻¹ ((coe_ideal_le_coe_ideal _).mpr h)) + (mul_left_mono (↑I)⁻¹ ((coe_ideal_le_coe_ideal _).mpr h)) (le_of_eq (inv_mul_cancel hI')), - obtain ⟨H, hH⟩ := fractional_ideal.le_one_iff_exists_coe_ideal.mp this, + obtain ⟨H, hH⟩ := le_one_iff_exists_coe_ideal.mp this, use H, refine coe_to_fractional_ideal_injective (le_refl (non_zero_divisors A)) (show (J : fractional_ideal A⁰ (fraction_ring A)) = _, from _), - rw [fractional_ideal.coe_ideal_mul, hH, ← mul_assoc, mul_inv_cancel hI', one_mul] + rw [coe_ideal_mul, hH, ← mul_assoc, mul_inv_cancel hI', one_mul] end⟩ lemma ideal.dvd_not_unit_iff_lt {I J : ideal A} : @@ -769,8 +753,8 @@ lemma ideal.exist_integer_multiples_not_mem ∃ i ∈ s, (a * f i) ∉ (J : fractional_ideal A⁰ K) := begin -- Consider the fractional ideal `I` spanned by the `f`s. - let I : fractional_ideal A⁰ K := fractional_ideal.span_finset A s f, - have hI0 : I ≠ 0 := fractional_ideal.span_finset_ne_zero.mpr ⟨j, hjs, hjf⟩, + let I : fractional_ideal A⁰ K := span_finset A s f, + have hI0 : I ≠ 0 := span_finset_ne_zero.mpr ⟨j, hjs, hjf⟩, -- We claim the multiplier `a` we're looking for is in `I⁻¹ \ (J / I)`. suffices : ↑J / I < I⁻¹, { obtain ⟨_, a, hI, hpI⟩ := set_like.lt_iff_le_and_exists.mp this, diff --git a/src/ring_theory/fractional_ideal.lean b/src/ring_theory/fractional_ideal.lean index 94189405de008..42cb8547169ec 100644 --- a/src/ring_theory/fractional_ideal.lean +++ b/src/ring_theory/fractional_ideal.lean @@ -500,8 +500,7 @@ begin end lemma coe_ideal_le_one {I : ideal R} : (I : fractional_ideal S P) ≤ 1 := -λ x hx, let ⟨y, _, hy⟩ := (fractional_ideal.mem_coe_ideal S).mp hx - in (fractional_ideal.mem_one_iff S).mpr ⟨y, hy⟩ +λ x hx, let ⟨y, _, hy⟩ := (mem_coe_ideal S).mp hx in (mem_one_iff S).mpr ⟨y, hy⟩ lemma le_one_iff_exists_coe_ideal {J : fractional_ideal S P} : J ≤ (1 : fractional_ideal S P) ↔ ∃ (I : ideal R), ↑I = J := @@ -522,7 +521,7 @@ begin { rintros ⟨y, hy, eq_y⟩, rwa ← eq_y }, { intro hx, - obtain ⟨y, eq_x⟩ := (fractional_ideal.mem_one_iff S).mp (hJ hx), + obtain ⟨y, eq_x⟩ := (mem_one_iff S).mp (hJ hx), rw ← eq_x at *, exact ⟨y, hx, rfl⟩ } } }, { rintro ⟨I, hI⟩, @@ -542,14 +541,14 @@ def coe_ideal_hom : ideal R →+* fractional_ideal S P := map_zero' := coe_to_fractional_ideal_bot } lemma coe_ideal_pow (I : ideal R) (n : ℕ) : (↑(I^n) : fractional_ideal S P) = I^n := -(fractional_ideal.coe_ideal_hom S P).map_pow _ n +(coe_ideal_hom S P).map_pow _ n open_locale big_operators lemma coe_ideal_finprod [is_localization S P] {α : Sort*} {f : α → ideal R} (hS : S ≤ non_zero_divisors R) : ((∏ᶠ a : α, f a : ideal R) : fractional_ideal S P) = ∏ᶠ a : α, (f a : fractional_ideal S P) := -monoid_hom.map_finprod_of_injective (fractional_ideal.coe_ideal_hom S P).to_monoid_hom - (fractional_ideal.coe_to_fractional_ideal_injective hS) f +monoid_hom.map_finprod_of_injective (coe_ideal_hom S P).to_monoid_hom + (coe_to_fractional_ideal_injective hS) f end order @@ -626,8 +625,7 @@ mem_map.trans ⟨λ ⟨x', hx', x'_eq⟩, h x'_eq ▸ hx', λ h, ⟨x, h, rfl⟩ lemma map_injective (f : P →ₐ[R] P') (h : function.injective f) : function.injective (map f : fractional_ideal S P → fractional_ideal S P') := -λ I J hIJ, fractional_ideal.ext (λ x, (fractional_ideal.map_mem_map h).symm.trans - (hIJ.symm ▸ fractional_ideal.map_mem_map h)) +λ I J hIJ, ext (λ x, (map_mem_map h).symm.trans (hIJ.symm ▸ map_mem_map h)) /-- If `g` is an equivalence, `map g` is an isomorphism -/ def map_equiv (g : P ≃ₐ[R] P') : @@ -695,8 +693,7 @@ begin obtain ⟨T, T', hT, hT', one_mem⟩ := mem_span_mul_finite_of_mem_mul this, refine ⟨T, submodule.span_eq_of_le _ hT _⟩, rw [← one_mul ↑I, ← mul_one (span R ↑T)], - conv_rhs { rw [← fractional_ideal.coe_one, ← units.mul_inv I, fractional_ideal.coe_mul, - mul_comm ↑↑I, ← mul_assoc] }, + conv_rhs { rw [← coe_one, ← units.mul_inv I, coe_mul, mul_comm ↑↑I, ← mul_assoc] }, refine submodule.mul_le_mul_left (le_trans _ (submodule.mul_le_mul_right (submodule.span_le.mpr hT'))), rwa [submodule.one_le, submodule.span_mul_span] @@ -1010,7 +1007,7 @@ lemma eq_zero_or_one (I : fractional_ideal K⁰ L) : I = 0 ∨ I = 1 := begin rw or_iff_not_imp_left, intro hI, - simp_rw [@set_like.ext_iff _ _ _ I 1, fractional_ideal.mem_one_iff], + simp_rw [@set_like.ext_iff _ _ _ I 1, mem_one_iff], intro x, split, { intro x_mem, @@ -1018,7 +1015,7 @@ begin refine ⟨n / d, _⟩, rw [map_div₀, is_fraction_ring.mk'_eq_div] }, { rintro ⟨x, rfl⟩, - obtain ⟨y, y_ne, y_mem⟩ := fractional_ideal.exists_ne_zero_mem_is_integer hI, + obtain ⟨y, y_ne, y_mem⟩ := exists_ne_zero_mem_is_integer hI, rw [← div_mul_cancel x y_ne, ring_hom.map_mul, ← algebra.smul_def], exact submodule.smul_mem I _ y_mem } end @@ -1038,8 +1035,7 @@ open_locale classical variables (R₁) /-- `fractional_ideal.span_finset R₁ s f` is the fractional ideal of `R₁` generated by `f '' s`. -/ -@[simps] def span_finset {ι : Type*} (s : finset ι) (f : ι → K) : - fractional_ideal R₁⁰ K := +@[simps] def span_finset {ι : Type*} (s : finset ι) (f : ι → K) : fractional_ideal R₁⁰ K := ⟨submodule.span R₁ (f '' s), begin obtain ⟨a', ha'⟩ := is_localization.exist_integer_multiples R₁⁰ s f, refine ⟨a', a'.2, λ x hx, submodule.span_induction hx _ _ _ _⟩, @@ -1201,8 +1197,7 @@ lemma mk'_mul_coe_ideal_eq_coe_ideal {I J : ideal R₁} {x y : R₁} (hy : y ∈ span_singleton R₁⁰ (is_localization.mk' K x ⟨y, hy⟩) * I = (J : fractional_ideal R₁⁰ K) ↔ ideal.span {x} * I = ideal.span {y} * J := begin - have inj : function.injective (coe : ideal R₁ → fractional_ideal R₁⁰ K) := - fractional_ideal.coe_ideal_injective, + have inj : function.injective (coe : ideal R₁ → fractional_ideal R₁⁰ K) := coe_ideal_injective, have : span_singleton R₁⁰ (is_localization.mk' _ (1 : R₁) ⟨y, hy⟩) * span_singleton R₁⁰ (algebra_map R₁ K y) = 1, { rw [span_singleton_mul_span_singleton, mul_comm, ← is_localization.mk'_eq_mul_mk'_one, @@ -1295,13 +1290,12 @@ include loc lemma le_span_singleton_mul_iff {x : P} {I J : fractional_ideal S P} : I ≤ span_singleton S x * J ↔ ∀ zI ∈ I, ∃ zJ ∈ J, x * zJ = zI := show (∀ {zI} (hzI : zI ∈ I), zI ∈ span_singleton _ x * J) ↔ ∀ zI ∈ I, ∃ zJ ∈ J, x * zJ = zI, -by simp only [fractional_ideal.mem_singleton_mul, eq_comm] +by simp only [mem_singleton_mul, eq_comm] lemma span_singleton_mul_le_iff {x : P} {I J : fractional_ideal S P} : span_singleton _ x * I ≤ J ↔ ∀ z ∈ I, x * z ∈ J := begin - simp only [fractional_ideal.mul_le, fractional_ideal.mem_singleton_mul, - fractional_ideal.mem_span_singleton], + simp only [mul_le, mem_singleton_mul, mem_span_singleton], split, { intros h zI hzI, exact h x ⟨1, one_smul _ _⟩ zI hzI }, @@ -1312,8 +1306,7 @@ end lemma eq_span_singleton_mul {x : P} {I J : fractional_ideal S P} : I = span_singleton _ x * J ↔ (∀ zI ∈ I, ∃ zJ ∈ J, x * zJ = zI) ∧ ∀ z ∈ J, x * z ∈ I := -by simp only [le_antisymm_iff, fractional_ideal.le_span_singleton_mul_iff, - fractional_ideal.span_singleton_mul_le_iff] +by simp only [le_antisymm_iff, le_span_singleton_mul_iff, span_singleton_mul_le_iff] end principal_ideal_ring From 9830a300340708eaa85d477c3fb96dd25f9468a5 Mon Sep 17 00:00:00 2001 From: Thomas Browning Date: Mon, 2 Jan 2023 14:48:24 +0000 Subject: [PATCH 0128/1291] chore(combinatorics/configuration): Generalize universes (#18039) This PR generalizes the universes, allowing the `P` (points) and `L` (lines) to have different universe levels. I also switched from forall to Pi for data fields, as suggested by @alreadydone. --- src/combinatorics/configuration.lean | 18 ++++++++---------- 1 file changed, 8 insertions(+), 10 deletions(-) diff --git a/src/combinatorics/configuration.lean b/src/combinatorics/configuration.lean index 045dc0b38b4bb..36221045516c0 100644 --- a/src/combinatorics/configuration.lean +++ b/src/combinatorics/configuration.lean @@ -36,9 +36,7 @@ open_locale big_operators namespace configuration -universe u - -variables (P L : Type u) [has_mem P L] +variables (P L : Type*) [has_mem P L] /-- A type synonym. -/ def dual := P @@ -63,13 +61,13 @@ class nondegenerate : Prop := (eq_or_eq : ∀ {p₁ p₂ : P} {l₁ l₂ : L}, p₁ ∈ l₁ → p₂ ∈ l₁ → p₁ ∈ l₂ → p₂ ∈ l₂ → p₁ = p₂ ∨ l₁ = l₂) /-- A nondegenerate configuration in which every pair of lines has an intersection point. -/ -class has_points extends nondegenerate P L : Type u := -(mk_point : ∀ {l₁ l₂ : L} (h : l₁ ≠ l₂), P) +class has_points extends nondegenerate P L := +(mk_point : Π {l₁ l₂ : L} (h : l₁ ≠ l₂), P) (mk_point_ax : ∀ {l₁ l₂ : L} (h : l₁ ≠ l₂), mk_point h ∈ l₁ ∧ mk_point h ∈ l₂) /-- A nondegenerate configuration in which every pair of points has a line through them. -/ -class has_lines extends nondegenerate P L : Type u := -(mk_line : ∀ {p₁ p₂ : P} (h : p₁ ≠ p₂), L) +class has_lines extends nondegenerate P L := +(mk_line : Π {p₁ p₂ : P} (h : p₁ ≠ p₂), L) (mk_line_ax : ∀ {p₁ p₂ : P} (h : p₁ ≠ p₂), p₁ ∈ mk_line h ∧ p₂ ∈ mk_line h) open nondegenerate has_points has_lines @@ -309,10 +307,10 @@ variables (P L) /-- A projective plane is a nondegenerate configuration in which every pair of lines has an intersection point, every pair of points has a line through them, and which has three points in general position. -/ -class projective_plane extends nondegenerate P L : Type u := -(mk_point : ∀ {l₁ l₂ : L} (h : l₁ ≠ l₂), P) +class projective_plane extends nondegenerate P L := +(mk_point : Π {l₁ l₂ : L} (h : l₁ ≠ l₂), P) (mk_point_ax : ∀ {l₁ l₂ : L} (h : l₁ ≠ l₂), mk_point h ∈ l₁ ∧ mk_point h ∈ l₂) -(mk_line : ∀ {p₁ p₂ : P} (h : p₁ ≠ p₂), L) +(mk_line : Π {p₁ p₂ : P} (h : p₁ ≠ p₂), L) (mk_line_ax : ∀ {p₁ p₂ : P} (h : p₁ ≠ p₂), p₁ ∈ mk_line h ∧ p₂ ∈ mk_line h) (exists_config : ∃ (p₁ p₂ p₃ : P) (l₁ l₂ l₃ : L), p₁ ∉ l₂ ∧ p₁ ∉ l₃ ∧ p₂ ∉ l₁ ∧ p₂ ∈ l₂ ∧ p₂ ∈ l₃ ∧ p₃ ∉ l₁ ∧ p₃ ∈ l₂ ∧ p₃ ∉ l₃) From 2da281a135d567ec86348cca09cc55f7ad9828c0 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 2 Jan 2023 21:17:23 +0000 Subject: [PATCH 0129/1291] fix(scripts/add_port_comments): do not add PR numbers to port comments (#18030) The now-automated port status YAML file no longer contains accurate PR numbers; many refer to ad-hoc ports rather than the real port. Since the PR numbers can no longer be trusted, we just omit them; both from the comment and the bot's commit message. --- .github/workflows/add_port_comment.yml | 7 +++---- scripts/add_port_comments.py | 3 +-- 2 files changed, 4 insertions(+), 6 deletions(-) diff --git a/.github/workflows/add_port_comment.yml b/.github/workflows/add_port_comment.yml index f57e4136b59a7..a65ed98ec6892 100644 --- a/.github/workflows/add_port_comment.yml +++ b/.github/workflows/add_port_comment.yml @@ -29,7 +29,7 @@ jobs: script: | proc = await exec.getExecOutput("python", ["./scripts/add_port_comments.py"]); console.log(proc.stdout); - core.setOutput("PR_LIST", proc.stdout); + core.setOutput("FILE_LIST", proc.stdout); - name: Create Pull Request uses: peter-evans/create-pull-request@v4 @@ -40,12 +40,11 @@ jobs: author: leanprover-community-bot body: | Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). - Relates to the following PRs: - ${{ steps.generate-message.outputs.PR_LIST }} + Relates to the following files: + ${{ steps.generate-message.outputs.FILE_LIST }} --- I am a bot; please check that I have not put a comment in a bad place before running `bors merge`! - If the PRs above don't match the file they are listed after, that means the port status page is wrong. labels: | easy awaiting-review diff --git a/scripts/add_port_comments.py b/scripts/add_port_comments.py index 35ef63c89ccb7..5c0160b5f9b56 100644 --- a/scripts/add_port_comments.py +++ b/scripts/add_port_comments.py @@ -16,7 +16,6 @@ def make_comment(fstatus): return textwrap.dedent(f"""\ > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. - > https://github.com/leanprover-community/mathlib4/pull/{fstatus.mathlib4_pr} > Any changes to this file require a corresponding PR to mathlib4.""") def replace_range(src: str, pos: int, end_pos: int, new: str) -> str: @@ -109,7 +108,7 @@ def fname_for(import_path: str) -> Path: continue if new_fcontent == fcontent: continue - print(f'* `{iname}`: https://github.com/leanprover-community/mathlib4/pull/{f_status.mathlib4_pr}') + print(f'* `{iname}`') with open(fname, 'w') as f: f.write(new_fcontent) if missing_docstrings: From 16de0889b9e841c59af6cfece272b9276f9bf5ae Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Mon, 2 Jan 2023 22:24:48 +0000 Subject: [PATCH 0130/1291] chore(*): golf (#18028) --- src/analysis/convex/specific_functions.lean | 14 ++++----- src/analysis/normed/field/basic.lean | 19 +++++------- src/number_theory/cyclotomic/basic.lean | 32 +++++++-------------- src/number_theory/prime_counting.lean | 2 +- 4 files changed, 25 insertions(+), 42 deletions(-) diff --git a/src/analysis/convex/specific_functions.lean b/src/analysis/convex/specific_functions.lean index 0b0513f6b21d6..a6a8429d91153 100644 --- a/src/analysis/convex/specific_functions.lean +++ b/src/analysis/convex/specific_functions.lean @@ -86,19 +86,17 @@ end lemma real.pow_sum_div_card_le_sum_pow {α : Type*} {s : finset α} {f : α → ℝ} (n : ℕ) (hf : ∀ a ∈ s, 0 ≤ f a) : (∑ x in s, f x) ^ (n + 1) / s.card ^ n ≤ ∑ x in s, (f x) ^ (n + 1) := begin - by_cases hs0 : s = ∅, - { simp_rw [hs0, finset.sum_empty, zero_pow' _ (nat.succ_ne_zero n), zero_div] }, - { have hs : s.card ≠ 0 := hs0 ∘ finset.card_eq_zero.1, - have hs' : (s.card : ℝ) ≠ 0 := (nat.cast_ne_zero.2 hs), - have hs'' : 0 < (s.card : ℝ) := nat.cast_pos.2 (nat.pos_of_ne_zero hs), + rcases s.eq_empty_or_nonempty with rfl | hs, + { simp_rw [finset.sum_empty, zero_pow' _ (nat.succ_ne_zero n), zero_div] }, + { have hs0 : 0 < (s.card : ℝ) := nat.cast_pos.2 hs.card_pos, suffices : (∑ x in s, f x / s.card) ^ (n + 1) ≤ ∑ x in s, (f x ^ (n + 1) / s.card), - by rwa [← finset.sum_div, ← finset.sum_div, div_pow, pow_succ' (s.card : ℝ), - ← div_div, div_le_iff hs'', div_mul, div_self hs', div_one] at this, + { rwa [← finset.sum_div, ← finset.sum_div, div_pow, pow_succ' (s.card : ℝ), + ← div_div, div_le_iff hs0, div_mul, div_self hs0.ne', div_one] at this }, have := @convex_on.map_sum_le ℝ ℝ ℝ α _ _ _ _ _ _ (set.Ici 0) (λ x, x ^ (n + 1)) s (λ _, 1 / s.card) (coe ∘ f) (convex_on_pow (n + 1)) _ _ (λ i hi, set.mem_Ici.2 (hf i hi)), { simpa only [inv_mul_eq_div, one_div, algebra.id.smul_eq_mul] using this }, { simp only [one_div, inv_nonneg, nat.cast_nonneg, implies_true_iff] }, - { simpa only [one_div, finset.sum_const, nsmul_eq_mul] using mul_inv_cancel hs' }} + { simpa only [one_div, finset.sum_const, nsmul_eq_mul] using mul_inv_cancel hs0.ne' } } end lemma nnreal.pow_sum_div_card_le_sum_pow {α : Type*} (s : finset α) (f : α → ℝ≥0) (n : ℕ) : diff --git a/src/analysis/normed/field/basic.lean b/src/analysis/normed/field/basic.lean index ae8b07693a29b..8e327edcd3906 100644 --- a/src/analysis/normed/field/basic.lean +++ b/src/analysis/normed/field/basic.lean @@ -489,23 +489,18 @@ begin simp, end -lemma norm_one_of_pow_eq_one {x : α} {k : ℕ+} (h : x ^ (k : ℕ) = 1) : - ‖x‖ = 1 := -begin - rw ( _ : ‖x‖ = 1 ↔ ‖x‖₊ = 1), - apply (@pow_left_inj nnreal _ _ _ ↑k zero_le' zero_le' (pnat.pos k)).mp, - { rw [← nnnorm_pow, one_pow, h, nnnorm_one], }, - { exact subtype.mk_eq_mk.symm, }, -end - -lemma norm_map_one_of_pow_eq_one [comm_monoid β] (φ : β →* α) {x : β} {k : ℕ+} +lemma norm_map_one_of_pow_eq_one [monoid β] (φ : β →* α) {x : β} {k : ℕ+} (h : x ^ (k : ℕ) = 1) : ‖φ x‖ = 1 := begin - have : (φ x) ^ (k : ℕ) = 1 := by rw [← monoid_hom.map_pow, h, monoid_hom.map_one], - exact norm_one_of_pow_eq_one this, + rw [← pow_left_inj, ← norm_pow, ← map_pow, h, map_one, norm_one, one_pow], + exacts [norm_nonneg _, zero_le_one, k.pos], end +lemma norm_one_of_pow_eq_one {x : α} {k : ℕ+} (h : x ^ (k : ℕ) = 1) : + ‖x‖ = 1 := +norm_map_one_of_pow_eq_one (monoid_hom.id α) h + end normed_division_ring /-- A normed field is a field with a norm satisfying ‖x y‖ = ‖x‖ ‖y‖. -/ diff --git a/src/number_theory/cyclotomic/basic.lean b/src/number_theory/cyclotomic/basic.lean index 027d406fad747..652614ab0c3fc 100644 --- a/src/number_theory/cyclotomic/basic.lean +++ b/src/number_theory/cyclotomic/basic.lean @@ -530,27 +530,17 @@ char_zero_of_injective_algebra_map ((algebra_map K _).injective) instance is_cyclotomic_extension [ne_zero ((n : ℕ) : K)] : is_cyclotomic_extension {n} K (cyclotomic_field n K) := -{ exists_prim_root := λ a han, - begin - rw mem_singleton_iff at han, - subst a, - obtain ⟨r, hr⟩ := exists_root_of_splits (algebra_map K (cyclotomic_field n K)) - (splitting_field.splits _) (degree_cyclotomic_pos n K (n.pos)).ne', - refine ⟨r, _⟩, - haveI := ne_zero.of_no_zero_smul_divisors K (cyclotomic_field n K) n, - rwa [← eval_map, ← is_root.def, map_cyclotomic, is_root_cyclotomic_iff] at hr - end, - adjoin_roots := - begin - rw [←algebra.eq_top_iff, ←splitting_field.adjoin_roots, eq_comm], - letI := classical.dec_eq (cyclotomic_field n K), - obtain ⟨ζ, hζ⟩ := exists_root_of_splits _ (splitting_field.splits (cyclotomic n K)) - (degree_cyclotomic_pos n _ n.pos).ne', - haveI : ne_zero ((n : ℕ) : (cyclotomic_field n K)) := - ne_zero.nat_of_injective (algebra_map K _).injective, - rw [eval₂_eq_eval_map, map_cyclotomic, ← is_root.def, is_root_cyclotomic_iff] at hζ, - exact is_cyclotomic_extension.adjoin_roots_cyclotomic_eq_adjoin_nth_roots hζ, - end } +begin + haveI : ne_zero ((n : ℕ) : (cyclotomic_field n K)) := + ne_zero.nat_of_injective (algebra_map K _).injective, + letI := classical.dec_eq (cyclotomic_field n K), + obtain ⟨ζ, hζ⟩ := exists_root_of_splits (algebra_map K (cyclotomic_field n K)) + (splitting_field.splits _) (degree_cyclotomic_pos n K n.pos).ne', + rw [← eval_map, ← is_root.def, map_cyclotomic, is_root_cyclotomic_iff] at hζ, + refine ⟨forall_eq.2 ⟨ζ, hζ⟩, _⟩, + rw [←algebra.eq_top_iff, ←splitting_field.adjoin_roots, eq_comm], + exact is_cyclotomic_extension.adjoin_roots_cyclotomic_eq_adjoin_nth_roots hζ, +end end cyclotomic_field diff --git a/src/number_theory/prime_counting.lean b/src/number_theory/prime_counting.lean index f36a914bc1ad7..b531b97019305 100644 --- a/src/number_theory/prime_counting.lean +++ b/src/number_theory/prime_counting.lean @@ -53,7 +53,7 @@ localized "notation (name := prime_counting') `π'` := nat.prime_counting'" in n lemma monotone_prime_counting' : monotone prime_counting' := count_monotone prime lemma monotone_prime_counting : monotone prime_counting := -λ a b a_le_b, monotone_prime_counting' (add_le_add_right a_le_b 1) +monotone_prime_counting'.comp (monotone_id.add_const _) @[simp] lemma prime_counting'_nth_eq (n : ℕ) : π' (nth prime n) = n := count_nth_of_infinite _ infinite_set_of_prime _ From cf73bb22be1b9070f789b9e51bd69951c4bfbb0e Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Mon, 2 Jan 2023 22:24:49 +0000 Subject: [PATCH 0131/1291] chore(number_theory/well_approximable): use nat.primes (#18042) --- src/number_theory/well_approximable.lean | 21 ++++++++++----------- 1 file changed, 10 insertions(+), 11 deletions(-) diff --git a/src/number_theory/well_approximable.lean b/src/number_theory/well_approximable.lean index c3bdd6abc385e..69255ac38e2c6 100644 --- a/src/number_theory/well_approximable.lean +++ b/src/number_theory/well_approximable.lean @@ -213,15 +213,14 @@ begin `E` is almost equal to `C p` for every prime. Combining this with 3 we find that `E` is almost invariant under the map `y ↦ y + 1/p` for every prime `p`. The required result then follows from `add_circle.ae_empty_or_univ_of_forall_vadd_ae_eq_self`. -/ + letI : semilattice_sup nat.primes := nat.subtype.semilattice_sup _, set μ : measure 𝕊 := volume, - set primes : set ℕ := {p | p.prime}, - haveI : nonempty primes := ⟨⟨2, nat.prime_two⟩⟩, - set u : primes → 𝕊 := λ p, ↑(((↑(1 : ℕ) : ℝ) / p) * T), - have hu₀ : ∀ (p : primes), add_order_of (u p) = (p : ℕ), + set u : nat.primes → 𝕊 := λ p, ↑(((↑(1 : ℕ) : ℝ) / p) * T), + have hu₀ : ∀ (p : nat.primes), add_order_of (u p) = (p : ℕ), { rintros ⟨p, hp⟩, exact add_order_of_div_of_gcd_eq_one hp.pos (gcd_one_left p), }, have hu : tendsto (add_order_of ∘ u) at_top at_top, { rw (funext hu₀ : add_order_of ∘ u = coe), - have h_mono : monotone (coe : primes → ℕ) := λ p q hpq, hpq, + have h_mono : monotone (coe : nat.primes → ℕ) := λ p q hpq, hpq, refine h_mono.tendsto_at_top_at_top (λ n, _), obtain ⟨p, hp, hp'⟩ := n.exists_infinite_primes, exact ⟨⟨p, hp'⟩, hp⟩, }, @@ -247,11 +246,11 @@ begin -- `tauto` can finish from here but unfortunately it's very slow. simp only [(em (p ∣ n)).symm, (em (p*p ∣ n)).symm, or_and_distrib_left, or_true, true_and, or_assoc], }, - have hE₂ : ∀ (p : primes), A p =ᵐ[μ] (∅ : set 𝕊) ∧ B p =ᵐ[μ] (∅ : set 𝕊) → E =ᵐ[μ] C p, + have hE₂ : ∀ (p : nat.primes), A p =ᵐ[μ] (∅ : set 𝕊) ∧ B p =ᵐ[μ] (∅ : set 𝕊) → E =ᵐ[μ] C p, { rintros p ⟨hA, hB⟩, rw hE₁ p, exact union_ae_eq_right_of_ae_eq_empty ((union_ae_eq_right_of_ae_eq_empty hA).trans hB), }, - have hA : ∀ (p : primes), A p =ᵐ[μ] (∅ : set 𝕊) ∨ A p =ᵐ[μ] univ, + have hA : ∀ (p : nat.primes), A p =ᵐ[μ] (∅ : set 𝕊) ∨ A p =ᵐ[μ] univ, { rintros ⟨p, hp⟩, let f : 𝕊 → 𝕊 := λ y, (p : ℕ) • y, suffices : f '' (A p) ⊆ @@ -263,7 +262,7 @@ begin refine (Sup_hom.set_image f).apply_blimsup_le.trans (mono_blimsup $ λ n hn, _), replace hn := nat.coprime_comm.mp (hp.coprime_iff_not_dvd.2 hn.2), exact approx_add_order_of.image_nsmul_subset_of_coprime (δ n) hp.pos hn, }, - have hB : ∀ (p : primes), B p =ᵐ[μ] (∅ : set 𝕊) ∨ B p =ᵐ[μ] univ, + have hB : ∀ (p : nat.primes), B p =ᵐ[μ] (∅ : set 𝕊) ∨ B p =ᵐ[μ] univ, { rintros ⟨p, hp⟩, let x := u ⟨p, hp⟩, let f : 𝕊 → 𝕊 := λ y, p • y + x, @@ -291,15 +290,15 @@ begin simp only [hu₀, subtype.coe_mk, h_div, mul_comm p], }, change (∀ᵐ x, x ∉ E) ∨ E ∈ volume.ae, rw [← eventually_eq_empty, ← eventually_eq_univ], - have hC : ∀ (p : primes), (u p) +ᵥ C p = C p, + have hC : ∀ (p : nat.primes), (u p) +ᵥ C p = C p, { intros p, let e := (add_action.to_perm (u p) : equiv.perm 𝕊).to_order_iso_set, change e (C p) = C p, rw [e.apply_blimsup, ← hu₀ p], exact blimsup_congr (eventually_of_forall $ λ n hn, approx_add_order_of.vadd_eq_of_mul_dvd (δ n) hn.1 hn.2), }, - by_cases h : ∀ (p : primes), A p =ᵐ[μ] (∅ : set 𝕊) ∧ B p =ᵐ[μ] (∅ : set 𝕊), - { replace h : ∀ (p : primes), ((u p) +ᵥ E : set _) =ᵐ[μ] E, + by_cases h : ∀ (p : nat.primes), A p =ᵐ[μ] (∅ : set 𝕊) ∧ B p =ᵐ[μ] (∅ : set 𝕊), + { replace h : ∀ (p : nat.primes), ((u p) +ᵥ E : set _) =ᵐ[μ] E, { intros p, replace hE₂ : E =ᵐ[μ] C p := hE₂ p (h p), have h_qmp : measure_theory.measure.quasi_measure_preserving ((+ᵥ) (-u p)) μ μ := From 2258b40dacd2942571c8ce136215350c702dc78f Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Mon, 2 Jan 2023 22:24:51 +0000 Subject: [PATCH 0132/1291] chore(analysis/inner_product_space/basic): golf proofs (#18043) --- src/analysis/inner_product_space/basic.lean | 16 ++++------------ 1 file changed, 4 insertions(+), 12 deletions(-) diff --git a/src/analysis/inner_product_space/basic.lean b/src/analysis/inner_product_space/basic.lean index 79de0faad1c92..fdac512a03076 100644 --- a/src/analysis/inner_product_space/basic.lean +++ b/src/analysis/inner_product_space/basic.lean @@ -943,11 +943,8 @@ end orthonormal_sets section norm lemma norm_eq_sqrt_inner (x : E) : ‖x‖ = sqrt (re ⟪x, x⟫) := -begin - have h₁ : ‖x‖^2 = re ⟪x, x⟫ := norm_sq_eq_inner x, - have h₂ := congr_arg sqrt h₁, - simpa only [sqrt_sq, norm_nonneg] using h₂, -end +calc ‖x‖ = sqrt (‖x‖ ^ 2) : (sqrt_sq (norm_nonneg _)).symm +... = sqrt (re ⟪x, x⟫) : congr_arg _ (norm_sq_eq_inner _) lemma norm_eq_sqrt_real_inner (x : F) : ‖x‖ = sqrt ⟪x, x⟫_ℝ := by { have h := @norm_eq_sqrt_inner ℝ F _ _ x, simpa using h } @@ -1554,13 +1551,8 @@ lemma abs_inner_eq_norm_iff (x y : E) (hx0 : x ≠ 0) (hy0 : y ≠ 0): begin have hxy0 : ‖x‖ * ‖y‖ ≠ 0 := mul_ne_zero (norm_eq_zero.not.2 hx0) (norm_eq_zero.not.2 hy0), have h₁ : abs ⟪x, y⟫ = ‖x‖ * ‖y‖ ↔ abs (⟪x, y⟫ / (‖x‖ * ‖y‖)) = 1, - { split; intro h, - { norm_cast, - rw [is_R_or_C.abs_div, h, abs_of_real, _root_.abs_mul, abs_norm_eq_norm, abs_norm_eq_norm], - exact div_self hxy0 }, - { norm_cast at h, - rwa [is_R_or_C.abs_div, abs_of_real, _root_.abs_mul, abs_norm_eq_norm, abs_norm_eq_norm, - div_eq_one_iff_eq hxy0] at h } }, + { rw [←algebra_map.coe_mul, is_R_or_C.abs_div, is_R_or_C.abs_of_nonneg, div_eq_one_iff_eq hxy0], + positivity }, rw [h₁, abs_inner_div_norm_mul_norm_eq_one_iff x y], exact and_iff_right hx0, end From c3291da49cfa65f0d43b094750541c0731edc932 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Tue, 3 Jan 2023 00:57:58 +0000 Subject: [PATCH 0133/1291] chore(*): add mathlib4 synchronization comments (#17979) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.associated` * `algebra.char_zero.lemmas` * `algebra.euclidean_domain.instances` * `algebra.field.opposite` * `algebra.field.power` * `algebra.group.conj` * `algebra.group.pi` * `algebra.group_power.lemmas` * `algebra.group_power.order` * `algebra.group_ring_action.basic` * `algebra.group_with_zero.power` * `algebra.hom.aut` * `algebra.hom.group_instances` * `algebra.hom.iterate` * `algebra.module.basic` * `algebra.module.prod` * `algebra.order.absolute_value` * `algebra.order.euclidean_absolute_value` * `algebra.order.field.basic` * `algebra.order.field.power` * `algebra.order.group.bounds` * `algebra.order.invertible` * `algebra.order.pi` * `algebra.order.positive.field` * `algebra.parity` * `algebra.regular.pow` * `algebra.regular.smul` * `algebra.ring.add_aut` * `algebra.ring.aut` * `algebra.ring.comp_typeclasses` * `algebra.ring.equiv` * `algebra.ring.opposite` * `algebra.ring.pi` * `algebra.ring.prod` * `algebra.ring.ulift` * `algebra.smul_with_zero` * `category_theory.category.Kleisli` * `category_theory.essential_image` * `category_theory.full_subcategory` * `category_theory.functor.fully_faithful` * `category_theory.whiskering` * `combinatorics.quiver.symmetric` * `control.fix` * `control.traversable.equiv` * `data.countable.small` * `data.erased` * `data.int.absolute_value` * `data.int.associated` * `data.int.bitwise` * `data.int.char_zero` * `data.int.conditionally_complete_order` * `data.int.dvd.pow` * `data.int.gcd` * `data.int.least_greatest` * `data.int.lemmas` * `data.int.modeq` * `data.int.nat_prime` * `data.int.order.lemmas` * `data.int.order.units` * `data.int.sqrt` * `data.list.func` * `data.nat.bits` * `data.nat.cast.field` * `data.nat.choose.basic` * `data.nat.choose.bounds` * `data.nat.choose.dvd` * `data.nat.dist` * `data.nat.even_odd_rec` * `data.nat.factorial.basic` * `data.nat.log` * `data.nat.modeq` * `data.nat.pairing` * `data.nat.pow` * `data.nat.prime` * `data.nat.size` * `data.nat.sqrt` * `data.nat.with_bot` * `data.part` * `data.pi.lex` * `data.pnat.basic` * `data.pnat.find` * `data.rat.basic` * `data.rat.defs` * `data.rat.lemmas` * `data.rat.order` * `data.rat.sqrt` * `data.rel` * `data.semiquot` * `data.set.Union_lift` * `data.set.accumulate` * `data.set.enumerate` * `data.set.function` * `data.set.functor` * `data.set.intervals.basic` * `data.set.intervals.disjoint` * `data.set.intervals.group` * `data.set.intervals.monoid` * `data.set.intervals.ord_connected` * `data.set.intervals.order_iso` * `data.set.intervals.pi` * `data.set.intervals.proj_Icc` * `data.set.intervals.surj_on` * `data.set.intervals.unordered_interval` * `data.set.intervals.with_bot_top` * `data.set.lattice` * `data.set.pairwise` * `data.setoid.basic` * `data.stream.init` * `dynamics.fixed_points.basic` * `group_theory.group_action.group` * `group_theory.group_action.opposite` * `group_theory.group_action.pi` * `group_theory.group_action.prod` * `group_theory.perm.basic` * `group_theory.perm.via_embedding` * `group_theory.submonoid.basic` * `group_theory.subsemigroup.basic` * `group_theory.subsemigroup.center` * `group_theory.subsemigroup.centralizer` * `group_theory.subsemigroup.operations` * `logic.equiv.nat` * `logic.equiv.set` * `logic.small.basic` * `order.antichain` * `order.atoms` * `order.bounded` * `order.bounds.basic` * `order.bounds.order_iso` * `order.chain` * `order.closure` * `order.complete_boolean_algebra` * `order.complete_lattice` * `order.complete_lattice_intervals` * `order.conditionally_complete_lattice.basic` * `order.conditionally_complete_lattice.group` * `order.copy` * `order.cover` * `order.fixed_points` * `order.galois_connection` * `order.heyting.regular` * `order.hom.order` * `order.hom.set` * `order.initial_seg` * `order.lattice_intervals` * `order.modular_lattice` * `order.monotone.extension` * `order.monotone.odd` * `order.monotone.union` * `order.ord_continuous` * `order.semiconj_Sup` * `order.succ_pred.basic` * `order.succ_pred.interval_succ` * `order.succ_pred.relation` * `order.zorn` * `order.zorn_atoms` * `ring_theory.ore_localization.ore_set` * `set_theory.cardinal.schroeder_bernstein` --- src/algebra/associated.lean | 3 +++ src/algebra/char_zero/lemmas.lean | 3 +++ src/algebra/euclidean_domain/instances.lean | 3 +++ src/algebra/field/opposite.lean | 3 +++ src/algebra/field/power.lean | 3 +++ src/algebra/group/conj.lean | 3 +++ src/algebra/group/pi.lean | 3 +++ src/algebra/group_power/lemmas.lean | 3 +++ src/algebra/group_power/order.lean | 3 +++ src/algebra/group_ring_action/basic.lean | 3 +++ src/algebra/group_with_zero/power.lean | 3 +++ src/algebra/hom/aut.lean | 3 +++ src/algebra/hom/group_instances.lean | 3 +++ src/algebra/hom/iterate.lean | 3 +++ src/algebra/module/basic.lean | 3 +++ src/algebra/module/prod.lean | 3 +++ src/algebra/order/absolute_value.lean | 3 +++ src/algebra/order/euclidean_absolute_value.lean | 3 +++ src/algebra/order/field/basic.lean | 3 +++ src/algebra/order/field/power.lean | 3 +++ src/algebra/order/group/bounds.lean | 3 +++ src/algebra/order/invertible.lean | 3 +++ src/algebra/order/pi.lean | 3 +++ src/algebra/order/positive/field.lean | 3 +++ src/algebra/parity.lean | 3 +++ src/algebra/regular/pow.lean | 3 +++ src/algebra/regular/smul.lean | 3 +++ src/algebra/ring/add_aut.lean | 3 +++ src/algebra/ring/aut.lean | 3 +++ src/algebra/ring/comp_typeclasses.lean | 3 +++ src/algebra/ring/equiv.lean | 3 +++ src/algebra/ring/opposite.lean | 3 +++ src/algebra/ring/pi.lean | 3 +++ src/algebra/ring/prod.lean | 3 +++ src/algebra/ring/ulift.lean | 3 +++ src/algebra/smul_with_zero.lean | 3 +++ src/category_theory/category/Kleisli.lean | 3 +++ src/category_theory/essential_image.lean | 3 +++ src/category_theory/full_subcategory.lean | 3 +++ src/category_theory/functor/fully_faithful.lean | 3 +++ src/category_theory/whiskering.lean | 3 +++ src/combinatorics/quiver/symmetric.lean | 3 +++ src/control/fix.lean | 3 +++ src/control/traversable/equiv.lean | 3 +++ src/data/countable/small.lean | 3 +++ src/data/erased.lean | 3 +++ src/data/int/absolute_value.lean | 3 +++ src/data/int/associated.lean | 3 +++ src/data/int/bitwise.lean | 3 +++ src/data/int/char_zero.lean | 3 +++ src/data/int/conditionally_complete_order.lean | 3 +++ src/data/int/dvd/pow.lean | 3 +++ src/data/int/gcd.lean | 3 +++ src/data/int/least_greatest.lean | 3 +++ src/data/int/lemmas.lean | 3 +++ src/data/int/modeq.lean | 3 +++ src/data/int/nat_prime.lean | 3 +++ src/data/int/order/lemmas.lean | 3 +++ src/data/int/order/units.lean | 3 +++ src/data/int/sqrt.lean | 3 +++ src/data/list/func.lean | 3 +++ src/data/nat/bits.lean | 3 +++ src/data/nat/cast/field.lean | 3 +++ src/data/nat/choose/basic.lean | 3 +++ src/data/nat/choose/bounds.lean | 3 +++ src/data/nat/choose/dvd.lean | 3 +++ src/data/nat/dist.lean | 3 +++ src/data/nat/even_odd_rec.lean | 5 ++++- src/data/nat/factorial/basic.lean | 3 +++ src/data/nat/log.lean | 3 +++ src/data/nat/modeq.lean | 3 +++ src/data/nat/pairing.lean | 3 +++ src/data/nat/pow.lean | 3 +++ src/data/nat/prime.lean | 3 +++ src/data/nat/size.lean | 5 ++++- src/data/nat/sqrt.lean | 3 +++ src/data/nat/with_bot.lean | 3 +++ src/data/part.lean | 3 +++ src/data/pi/lex.lean | 3 +++ src/data/pnat/basic.lean | 3 +++ src/data/pnat/find.lean | 3 +++ src/data/rat/basic.lean | 3 +++ src/data/rat/defs.lean | 3 +++ src/data/rat/lemmas.lean | 3 +++ src/data/rat/order.lean | 3 +++ src/data/rat/sqrt.lean | 3 +++ src/data/rel.lean | 3 +++ src/data/semiquot.lean | 3 +++ src/data/set/Union_lift.lean | 3 +++ src/data/set/accumulate.lean | 3 +++ src/data/set/enumerate.lean | 3 +++ src/data/set/function.lean | 3 +++ src/data/set/functor.lean | 3 +++ src/data/set/intervals/basic.lean | 3 +++ src/data/set/intervals/disjoint.lean | 3 +++ src/data/set/intervals/group.lean | 5 ++++- src/data/set/intervals/monoid.lean | 3 +++ src/data/set/intervals/ord_connected.lean | 3 +++ src/data/set/intervals/order_iso.lean | 3 +++ src/data/set/intervals/pi.lean | 3 +++ src/data/set/intervals/proj_Icc.lean | 3 +++ src/data/set/intervals/surj_on.lean | 3 +++ src/data/set/intervals/unordered_interval.lean | 3 +++ src/data/set/intervals/with_bot_top.lean | 3 +++ src/data/set/lattice.lean | 3 +++ src/data/set/pairwise.lean | 3 +++ src/data/setoid/basic.lean | 3 +++ src/data/stream/init.lean | 3 +++ src/dynamics/fixed_points/basic.lean | 3 +++ src/group_theory/group_action/group.lean | 3 +++ src/group_theory/group_action/opposite.lean | 3 +++ src/group_theory/group_action/pi.lean | 3 +++ src/group_theory/group_action/prod.lean | 3 +++ src/group_theory/perm/basic.lean | 3 +++ src/group_theory/perm/via_embedding.lean | 3 +++ src/group_theory/submonoid/basic.lean | 3 +++ src/group_theory/subsemigroup/basic.lean | 3 +++ src/group_theory/subsemigroup/center.lean | 3 +++ src/group_theory/subsemigroup/centralizer.lean | 3 +++ src/group_theory/subsemigroup/operations.lean | 3 +++ src/logic/equiv/nat.lean | 3 +++ src/logic/equiv/set.lean | 3 +++ src/logic/small/basic.lean | 3 +++ src/order/antichain.lean | 3 +++ src/order/atoms.lean | 3 +++ src/order/bounded.lean | 3 +++ src/order/bounds/basic.lean | 3 +++ src/order/bounds/order_iso.lean | 3 +++ src/order/chain.lean | 3 +++ src/order/closure.lean | 3 +++ src/order/complete_boolean_algebra.lean | 3 +++ src/order/complete_lattice.lean | 3 +++ src/order/complete_lattice_intervals.lean | 3 +++ src/order/conditionally_complete_lattice/basic.lean | 3 +++ src/order/conditionally_complete_lattice/group.lean | 3 +++ src/order/copy.lean | 3 +++ src/order/cover.lean | 3 +++ src/order/fixed_points.lean | 3 +++ src/order/galois_connection.lean | 3 +++ src/order/heyting/regular.lean | 3 +++ src/order/hom/order.lean | 3 +++ src/order/hom/set.lean | 3 +++ src/order/initial_seg.lean | 3 +++ src/order/lattice_intervals.lean | 3 +++ src/order/modular_lattice.lean | 3 +++ src/order/monotone/extension.lean | 3 +++ src/order/monotone/odd.lean | 3 +++ src/order/monotone/union.lean | 3 +++ src/order/ord_continuous.lean | 3 +++ src/order/semiconj_Sup.lean | 3 +++ src/order/succ_pred/basic.lean | 3 +++ src/order/succ_pred/interval_succ.lean | 3 +++ src/order/succ_pred/relation.lean | 3 +++ src/order/zorn.lean | 3 +++ src/order/zorn_atoms.lean | 3 +++ src/ring_theory/ore_localization/ore_set.lean | 3 +++ src/set_theory/cardinal/schroeder_bernstein.lean | 3 +++ 157 files changed, 474 insertions(+), 3 deletions(-) diff --git a/src/algebra/associated.lean b/src/algebra/associated.lean index b955e8b4be8e8..a77742482a09d 100644 --- a/src/algebra/associated.lean +++ b/src/algebra/associated.lean @@ -9,6 +9,9 @@ import algebra.parity /-! # Associated, prime, and irreducible elements. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ variables {α : Type*} {β : Type*} {γ : Type*} {δ : Type*} diff --git a/src/algebra/char_zero/lemmas.lean b/src/algebra/char_zero/lemmas.lean index d7defafec64dd..ea21ced81e110 100644 --- a/src/algebra/char_zero/lemmas.lean +++ b/src/algebra/char_zero/lemmas.lean @@ -10,6 +10,9 @@ import algebra.group_power.lemmas /-! # Characteristic zero (additional theorems) +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A ring `R` is called of characteristic zero if every natural number `n` is non-zero when considered as an element of `R`. Since this definition doesn't mention the multiplicative structure of `R` except for the existence of `1` in this file characteristic zero is defined for additive monoids diff --git a/src/algebra/euclidean_domain/instances.lean b/src/algebra/euclidean_domain/instances.lean index 3d1841412230d..59e62167a406d 100644 --- a/src/algebra/euclidean_domain/instances.lean +++ b/src/algebra/euclidean_domain/instances.lean @@ -12,6 +12,9 @@ import data.int.order.basic /-! # Instances for Euclidean domains +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + * `int.euclidean_domain`: shows that `ℤ` is a Euclidean domain. * `field.to_euclidean_domain`: shows that any field is a Euclidean domain. diff --git a/src/algebra/field/opposite.lean b/src/algebra/field/opposite.lean index 89d4f2b268d42..b7c310ef76717 100644 --- a/src/algebra/field/opposite.lean +++ b/src/algebra/field/opposite.lean @@ -8,6 +8,9 @@ import algebra.ring.opposite /-! # Field structure on the multiplicative/additive opposite + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ variables (α : Type*) diff --git a/src/algebra/field/power.lean b/src/algebra/field/power.lean index 502aba88feb93..b02730a6048fb 100644 --- a/src/algebra/field/power.lean +++ b/src/algebra/field/power.lean @@ -10,6 +10,9 @@ import algebra.parity /-! # Results about powers in fields or division rings. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file exists to ensure we can define `field` with minimal imports, so contains some lemmas about powers of elements which need imports beyond those needed for the basic definition. diff --git a/src/algebra/group/conj.lean b/src/algebra/group/conj.lean index 6912b20fa5912..ba988ce048bf5 100644 --- a/src/algebra/group/conj.lean +++ b/src/algebra/group/conj.lean @@ -11,6 +11,9 @@ import algebra.hom.group /-! # Conjugacy of group elements +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + See also `mul_aut.conj` and `quandle.conj`. -/ diff --git a/src/algebra/group/pi.lean b/src/algebra/group/pi.lean index 52c2eb9a86a82..0039f46c13aba 100644 --- a/src/algebra/group/pi.lean +++ b/src/algebra/group/pi.lean @@ -12,6 +12,9 @@ import tactic.pi_instances /-! # Pi instances for groups and monoids +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines instances for group, monoid, semigroup and related structures on Pi types. -/ diff --git a/src/algebra/group_power/lemmas.lean b/src/algebra/group_power/lemmas.lean index aeb7915b1fad1..39cb2da938b71 100644 --- a/src/algebra/group_power/lemmas.lean +++ b/src/algebra/group_power/lemmas.lean @@ -12,6 +12,9 @@ import data.int.cast.lemmas /-! # Lemmas about power operations on monoids and groups +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains lemmas about `monoid.pow`, `group.pow`, `nsmul`, `zsmul` which require additional imports besides those available in `algebra.group_power.basic`. -/ diff --git a/src/algebra/group_power/order.lean b/src/algebra/group_power/order.lean index cd935027fefcb..773cadffb22e5 100644 --- a/src/algebra/group_power/order.lean +++ b/src/algebra/group_power/order.lean @@ -11,6 +11,9 @@ import data.set.intervals.basic /-! # Lemmas about the interaction of power operations with order +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Note that some lemmas are in `algebra/group_power/lemmas.lean` as they import files which depend on this file. -/ diff --git a/src/algebra/group_ring_action/basic.lean b/src/algebra/group_ring_action/basic.lean index 539cca73b344f..c51cd994effd6 100644 --- a/src/algebra/group_ring_action/basic.lean +++ b/src/algebra/group_ring_action/basic.lean @@ -11,6 +11,9 @@ import group_theory.group_action.group /-! # Group action on rings +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the typeclass of monoid acting on semirings `mul_semiring_action M R`, and the corresponding typeclass of invariant subrings. diff --git a/src/algebra/group_with_zero/power.lean b/src/algebra/group_with_zero/power.lean index ffd1741956fac..04129dafa0dc6 100644 --- a/src/algebra/group_with_zero/power.lean +++ b/src/algebra/group_with_zero/power.lean @@ -9,6 +9,9 @@ import data.int.bitwise /-! # Powers of elements of groups with an adjoined zero element +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define integer power functions for groups with an adjoined zero element. This generalises the integer power function on a division ring. -/ diff --git a/src/algebra/hom/aut.lean b/src/algebra/hom/aut.lean index f48ac064ff47b..d0e2bfe47879f 100644 --- a/src/algebra/hom/aut.lean +++ b/src/algebra/hom/aut.lean @@ -8,6 +8,9 @@ import group_theory.perm.basic /-! # Multiplicative and additive group automorphisms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the automorphism group structure on `add_aut R := add_equiv R R` and `mul_aut R := mul_equiv R R`. diff --git a/src/algebra/hom/group_instances.lean b/src/algebra/hom/group_instances.lean index 9503382201421..ba5ff959d7466 100644 --- a/src/algebra/hom/group_instances.lean +++ b/src/algebra/hom/group_instances.lean @@ -10,6 +10,9 @@ import algebra.ring.basic /-! # Instances on spaces of monoid and group morphisms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We endow the space of monoid morphisms `M →* N` with a `comm_monoid` structure when the target is commutative, through pointwise multiplication, and with a `comm_group` structure when the target is a commutative group. We also prove the same instances for additive situations. diff --git a/src/algebra/hom/iterate.lean b/src/algebra/hom/iterate.lean index 726af4b26357a..b9e900512f21a 100644 --- a/src/algebra/hom/iterate.lean +++ b/src/algebra/hom/iterate.lean @@ -11,6 +11,9 @@ import group_theory.group_action.opposite /-! # Iterates of monoid and ring homomorphisms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Iterate of a monoid/ring homomorphism is a monoid/ring homomorphism but it has a wrong type, so Lean can't apply lemmas like `monoid_hom.map_one` to `f^[n] 1`. Though it is possible to define a monoid structure on the endomorphisms, quite often we do not want to convert from diff --git a/src/algebra/module/basic.lean b/src/algebra/module/basic.lean index 4aa56282af5f3..a242e541adc8a 100644 --- a/src/algebra/module/basic.lean +++ b/src/algebra/module/basic.lean @@ -11,6 +11,9 @@ import tactic.abel /-! # Modules over a ring +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define * `module R M` : an additive commutative monoid `M` is a `module` over a diff --git a/src/algebra/module/prod.lean b/src/algebra/module/prod.lean index 020eb4accd320..3f5d9ca458169 100644 --- a/src/algebra/module/prod.lean +++ b/src/algebra/module/prod.lean @@ -9,6 +9,9 @@ import group_theory.group_action.prod /-! # Prod instances for module and multiplicative actions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines instances for binary product of modules -/ diff --git a/src/algebra/order/absolute_value.lean b/src/algebra/order/absolute_value.lean index fa3cf5f55fc3b..adfde39ee76c6 100644 --- a/src/algebra/order/absolute_value.lean +++ b/src/algebra/order/absolute_value.lean @@ -13,6 +13,9 @@ import algebra.ring.regular /-! # Absolute values +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines a bundled type of absolute values `absolute_value R S`. ## Main definitions diff --git a/src/algebra/order/euclidean_absolute_value.lean b/src/algebra/order/euclidean_absolute_value.lean index 913d36d4b278b..cef698d857b78 100644 --- a/src/algebra/order/euclidean_absolute_value.lean +++ b/src/algebra/order/euclidean_absolute_value.lean @@ -9,6 +9,9 @@ import algebra.euclidean_domain.instances /-! # Euclidean absolute values +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines a predicate `absolute_value.is_euclidean abv` stating the absolute value is compatible with the Euclidean domain structure on its domain. diff --git a/src/algebra/order/field/basic.lean b/src/algebra/order/field/basic.lean index 8318abff932d4..f7a94a50a75d3 100644 --- a/src/algebra/order/field/basic.lean +++ b/src/algebra/order/field/basic.lean @@ -11,6 +11,9 @@ import algebra.group_power.order /-! # Lemmas about linear ordered (semi)fields +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + -/ open function order_dual diff --git a/src/algebra/order/field/power.lean b/src/algebra/order/field/power.lean index aa69291ff6968..8c265dd4490bd 100644 --- a/src/algebra/order/field/power.lean +++ b/src/algebra/order/field/power.lean @@ -10,6 +10,9 @@ import algebra.order.field.basic /-! # Lemmas about powers in ordered fields. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ variables {α : Type*} diff --git a/src/algebra/order/group/bounds.lean b/src/algebra/order/group/bounds.lean index 85767cd6b059e..13573ec0b6178 100644 --- a/src/algebra/order/group/bounds.lean +++ b/src/algebra/order/group/bounds.lean @@ -8,6 +8,9 @@ import algebra.order.group.defs /-! # Least upper bound and the greatest lower bound in linear ordered additive commutative groups + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ variables {α : Type*} diff --git a/src/algebra/order/invertible.lean b/src/algebra/order/invertible.lean index d862f4ae21ef8..a6d6d8b08736c 100644 --- a/src/algebra/order/invertible.lean +++ b/src/algebra/order/invertible.lean @@ -8,6 +8,9 @@ import algebra.invertible /-! # Lemmas about `inv_of` in ordered (semi)rings. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ variables {α : Type*} [linear_ordered_semiring α] {a : α} diff --git a/src/algebra/order/pi.lean b/src/algebra/order/pi.lean index bc62e3627bcc2..7c696ad7209e0 100644 --- a/src/algebra/order/pi.lean +++ b/src/algebra/order/pi.lean @@ -10,6 +10,9 @@ import tactic.positivity /-! # Pi instances for ordered groups and monoids +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines instances for ordered group, monoid, and related structures on Pi types. -/ diff --git a/src/algebra/order/positive/field.lean b/src/algebra/order/positive/field.lean index 1f17bb920c68b..c5a955471ca01 100644 --- a/src/algebra/order/positive/field.lean +++ b/src/algebra/order/positive/field.lean @@ -9,6 +9,9 @@ import algebra.order.positive.ring /-! # Algebraic structures on the set of positive numbers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove that the set of positive elements of a linear ordered field is a linear ordered commutative group. -/ diff --git a/src/algebra/parity.lean b/src/algebra/parity.lean index 61410fc5820ec..e57fc2b87e6a5 100644 --- a/src/algebra/parity.lean +++ b/src/algebra/parity.lean @@ -7,6 +7,9 @@ import algebra.group_power.lemmas /-! # Squares, even and odd elements +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves some general facts about squares, even and odd elements of semirings. In the implementation, we define `is_square` and we let `even` be the notion transported by diff --git a/src/algebra/regular/pow.lean b/src/algebra/regular/pow.lean index 696e8fdbf250c..ee2d02d7819ad 100644 --- a/src/algebra/regular/pow.lean +++ b/src/algebra/regular/pow.lean @@ -9,6 +9,9 @@ import algebra.regular.basic /-! # Regular elements +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Implementation details Group powers and other definitions import a lot of the algebra hierarchy. diff --git a/src/algebra/regular/smul.lean b/src/algebra/regular/smul.lean index be04083079e37..6ec31dc4b434b 100644 --- a/src/algebra/regular/smul.lean +++ b/src/algebra/regular/smul.lean @@ -9,6 +9,9 @@ import algebra.regular.basic /-! # Action of regular elements on a module +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We introduce `M`-regular elements, in the context of an `R`-module `M`. The corresponding predicate is called `is_smul_regular`. diff --git a/src/algebra/ring/add_aut.lean b/src/algebra/ring/add_aut.lean index f92c57847c04a..986c4711c60c6 100644 --- a/src/algebra/ring/add_aut.lean +++ b/src/algebra/ring/add_aut.lean @@ -9,6 +9,9 @@ import algebra.module.basic /-! # Multiplication on the left/right as additive automorphisms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define `add_aut.mul_left` and `add_aut.mul_right`. See also `add_monoid_hom.mul_left`, `add_monoid_hom.mul_right`, `add_monoid.End.mul_left`, and diff --git a/src/algebra/ring/aut.lean b/src/algebra/ring/aut.lean index 2d77f91182fd2..05731948231e1 100644 --- a/src/algebra/ring/aut.lean +++ b/src/algebra/ring/aut.lean @@ -10,6 +10,9 @@ import algebra.ring.equiv /-! # Ring automorphisms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the automorphism group structure on `ring_aut R := ring_equiv R R`. ## Implementation notes diff --git a/src/algebra/ring/comp_typeclasses.lean b/src/algebra/ring/comp_typeclasses.lean index 7ab5690049c69..19f4fc328812e 100644 --- a/src/algebra/ring/comp_typeclasses.lean +++ b/src/algebra/ring/comp_typeclasses.lean @@ -9,6 +9,9 @@ import algebra.ring.equiv /-! # Propositional typeclasses on several ring homs +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains three typeclasses used in the definition of (semi)linear maps: * `ring_hom_comp_triple σ₁₂ σ₂₃ σ₁₃`, which expresses the fact that `σ₂₃.comp σ₁₂ = σ₁₃` * `ring_hom_inv_pair σ₁₂ σ₂₁`, which states that `σ₁₂` and `σ₂₁` are inverses of each other diff --git a/src/algebra/ring/equiv.lean b/src/algebra/ring/equiv.lean index f5686ae575014..224d8eb3c521f 100644 --- a/src/algebra/ring/equiv.lean +++ b/src/algebra/ring/equiv.lean @@ -11,6 +11,9 @@ import tactic.assert_exists /-! # (Semi)ring equivs +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define extension of `equiv` called `ring_equiv`, which is a datatype representing an isomorphism of `semiring`s, `ring`s, `division_ring`s, or `field`s. We also introduce the corresponding group of automorphisms `ring_aut`. diff --git a/src/algebra/ring/opposite.lean b/src/algebra/ring/opposite.lean index 55be00f3ffae3..e9d57dd7a7955 100644 --- a/src/algebra/ring/opposite.lean +++ b/src/algebra/ring/opposite.lean @@ -9,6 +9,9 @@ import algebra.hom.ring /-! # Ring structures on the multiplicative opposite + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ universes u v variables (α : Type u) diff --git a/src/algebra/ring/pi.lean b/src/algebra/ring/pi.lean index 380d0d8d4b040..b56dbf7fafd41 100644 --- a/src/algebra/ring/pi.lean +++ b/src/algebra/ring/pi.lean @@ -10,6 +10,9 @@ import algebra.hom.ring /-! # Pi instances for ring +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines instances for ring, semiring and related structures on Pi Types -/ diff --git a/src/algebra/ring/prod.lean b/src/algebra/ring/prod.lean index a92376d1a0f37..1ac110c8c81a5 100644 --- a/src/algebra/ring/prod.lean +++ b/src/algebra/ring/prod.lean @@ -11,6 +11,9 @@ import algebra.order.monoid.prod /-! # Semiring, ring etc structures on `R × S` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define two-binop (`semiring`, `ring` etc) structures on `R × S`. We also prove trivial `simp` lemmas, and define the following operations on `ring_hom`s and similarly for `non_unital_ring_hom`s: diff --git a/src/algebra/ring/ulift.lean b/src/algebra/ring/ulift.lean index c16e7a9fcac81..6336fa3f6fe76 100644 --- a/src/algebra/ring/ulift.lean +++ b/src/algebra/ring/ulift.lean @@ -10,6 +10,9 @@ import algebra.ring.equiv /-! # `ulift` instances for ring +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines instances for ring, semiring and related structures on `ulift` types. (Recall `ulift α` is just a "copy" of a type `α` in a higher universe.) diff --git a/src/algebra/smul_with_zero.lean b/src/algebra/smul_with_zero.lean index 6a0c9db627f03..fd32bf48762f0 100644 --- a/src/algebra/smul_with_zero.lean +++ b/src/algebra/smul_with_zero.lean @@ -11,6 +11,9 @@ import group_theory.group_action.prod /-! # Introduce `smul_with_zero` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In analogy with the usual monoid action on a Type `M`, we introduce an action of a `monoid_with_zero` on a Type with `0`. diff --git a/src/category_theory/category/Kleisli.lean b/src/category_theory/category/Kleisli.lean index 2dd2d7d5c431a..2c5d47ad9e81c 100644 --- a/src/category_theory/category/Kleisli.lean +++ b/src/category_theory/category/Kleisli.lean @@ -8,6 +8,9 @@ import category_theory.category.basic /-! # The Kleisli construction on the Type category +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Define the Kleisli category for (control) monads. `category_theory/monad/kleisli` defines the general version for a monad on `C`, and demonstrates the equivalence between the two. diff --git a/src/category_theory/essential_image.lean b/src/category_theory/essential_image.lean index 374ac99ad2a49..c37336e2295e0 100644 --- a/src/category_theory/essential_image.lean +++ b/src/category_theory/essential_image.lean @@ -9,6 +9,9 @@ import category_theory.full_subcategory /-! # Essential image of a functor +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The essential image `ess_image` of a functor consists of the objects in the target category which are isomorphic to an object in the image of the object function. This, for instance, allows us to talk about objects belonging to a subcategory expressed as a diff --git a/src/category_theory/full_subcategory.lean b/src/category_theory/full_subcategory.lean index 52397e39da4f6..e80c647354ae7 100644 --- a/src/category_theory/full_subcategory.lean +++ b/src/category_theory/full_subcategory.lean @@ -8,6 +8,9 @@ import category_theory.functor.fully_faithful /-! # Induced categories and full subcategories +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given a category `D` and a function `F : C → D `from a type `C` to the objects of `D`, there is an essentially unique way to give `C` a category structure such that `F` becomes a fully faithful functor, diff --git a/src/category_theory/functor/fully_faithful.lean b/src/category_theory/functor/fully_faithful.lean index 400658ba2ff00..e62ccf3d84f75 100644 --- a/src/category_theory/functor/fully_faithful.lean +++ b/src/category_theory/functor/fully_faithful.lean @@ -9,6 +9,9 @@ import logic.equiv.defs /-! # Full and faithful functors +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define typeclasses `full` and `faithful`, decorating functors. ## Main definitions and results diff --git a/src/category_theory/whiskering.lean b/src/category_theory/whiskering.lean index b50448b1a49db..60bed39c4592b 100644 --- a/src/category_theory/whiskering.lean +++ b/src/category_theory/whiskering.lean @@ -10,6 +10,9 @@ import category_theory.functor.fully_faithful /-! # Whiskering +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given a functor `F : C ⥤ D` and functors `G H : D ⥤ E` and a natural transformation `α : G ⟶ H`, we can construct a new natural transformation `F ⋙ G ⟶ F ⋙ H`, called `whisker_left F α`. This is the same as the horizontal composition of `𝟙 F` with `α`. diff --git a/src/combinatorics/quiver/symmetric.lean b/src/combinatorics/quiver/symmetric.lean index 156867f50782e..640d4b1005b52 100644 --- a/src/combinatorics/quiver/symmetric.lean +++ b/src/combinatorics/quiver/symmetric.lean @@ -10,6 +10,9 @@ import data.sum.basic /-! ## Symmetric quivers and arrow reversal +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains constructions related to symmetric quivers: * `symmetrify V` adds formal inverses to each arrow of `V`. diff --git a/src/control/fix.lean b/src/control/fix.lean index 326c5daa2ac42..aae0456ab23e6 100644 --- a/src/control/fix.lean +++ b/src/control/fix.lean @@ -11,6 +11,9 @@ import data.stream.defs /-! # Fixed point +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This module defines a generic `fix` operator for defining recursive computations that are not necessarily well-founded or productive. An instance is defined for `part`. diff --git a/src/control/traversable/equiv.lean b/src/control/traversable/equiv.lean index ea9d1588094a3..205a4a8d4d7aa 100644 --- a/src/control/traversable/equiv.lean +++ b/src/control/traversable/equiv.lean @@ -9,6 +9,9 @@ import logic.equiv.defs /-! # Transferring `traversable` instances along isomorphisms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file allows to transfer `traversable` instances along isomorphisms. ## Main declarations diff --git a/src/data/countable/small.lean b/src/data/countable/small.lean index 95d88c7fb94b3..6b2280e7188fb 100644 --- a/src/data/countable/small.lean +++ b/src/data/countable/small.lean @@ -9,6 +9,9 @@ import data.countable.defs /-! # All countable types are small. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + That is, any countable type is equivalent to a type in any universe. -/ diff --git a/src/data/erased.lean b/src/data/erased.lean index b9ffd5bbde2eb..6831006f15927 100644 --- a/src/data/erased.lean +++ b/src/data/erased.lean @@ -8,6 +8,9 @@ import logic.equiv.defs /-! # A type for VM-erased data +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines a type `erased α` which is classically isomorphic to `α`, but erased in the VM. That is, at runtime every value of `erased α` is represented as `0`, just like types and proofs. diff --git a/src/data/int/absolute_value.lean b/src/data/int/absolute_value.lean index 745c78c96b006..754cf1b10bbb0 100644 --- a/src/data/int/absolute_value.lean +++ b/src/data/int/absolute_value.lean @@ -12,6 +12,9 @@ import group_theory.group_action.units /-! # Absolute values and the integers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains some results on absolute values applied to integers. ## Main results diff --git a/src/data/int/associated.lean b/src/data/int/associated.lean index 0927f1b24c1f7..1aae01eef87ac 100644 --- a/src/data/int/associated.lean +++ b/src/data/int/associated.lean @@ -10,6 +10,9 @@ import data.int.units /-! # Associated elements and the integers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains some results on equality up to units in the integers. ## Main results diff --git a/src/data/int/bitwise.lean b/src/data/int/bitwise.lean index 73fb86f91588f..0114fb8fc132e 100644 --- a/src/data/int/bitwise.lean +++ b/src/data/int/bitwise.lean @@ -10,6 +10,9 @@ import data.nat.size /-! # Bitwise operations on integers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Recursors * `int.bit_cases_on`: Parity disjunction. Something is true/defined on `ℤ` if it's true/defined for diff --git a/src/data/int/char_zero.lean b/src/data/int/char_zero.lean index 3ce54c05a6f57..8a6cb00890c84 100644 --- a/src/data/int/char_zero.lean +++ b/src/data/int/char_zero.lean @@ -8,6 +8,9 @@ import data.int.cast.field /-! # Injectivity of `int.cast` into characteristic zero rings and fields. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + -/ variables {α : Type*} diff --git a/src/data/int/conditionally_complete_order.lean b/src/data/int/conditionally_complete_order.lean index d1a7e50388e3a..63efc9885b7e2 100644 --- a/src/data/int/conditionally_complete_order.lean +++ b/src/data/int/conditionally_complete_order.lean @@ -9,6 +9,9 @@ import data.int.least_greatest /-! ## `ℤ` forms a conditionally complete linear order +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The integers form a conditionally complete linear order. -/ diff --git a/src/data/int/dvd/pow.lean b/src/data/int/dvd/pow.lean index e3e6442322d02..1c74d4e8d51a7 100644 --- a/src/data/int/dvd/pow.lean +++ b/src/data/int/dvd/pow.lean @@ -8,6 +8,9 @@ import data.nat.pow /-! # Basic lemmas about the divisibility relation in `ℤ` involving powers. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ open nat diff --git a/src/data/int/gcd.lean b/src/data/int/gcd.lean index de63a730246ff..50f50b40dd6b5 100644 --- a/src/data/int/gcd.lean +++ b/src/data/int/gcd.lean @@ -9,6 +9,9 @@ import tactic.norm_num /-! # Extended GCD and divisibility over ℤ +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions * Given `x y : ℕ`, `xgcd x y` computes the pair of integers `(a, b)` such that diff --git a/src/data/int/least_greatest.lean b/src/data/int/least_greatest.lean index ae103739612d0..f9d07dcc6c83d 100644 --- a/src/data/int/least_greatest.lean +++ b/src/data/int/least_greatest.lean @@ -7,6 +7,9 @@ import data.int.order.basic /-! # Least upper bound and greatest lower bound properties for integers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove that a bounded above nonempty set of integers has the greatest element, and a counterpart of this statement for the least element. diff --git a/src/data/int/lemmas.lean b/src/data/int/lemmas.lean index e5effced31ce8..f3192a3f8d971 100644 --- a/src/data/int/lemmas.lean +++ b/src/data/int/lemmas.lean @@ -12,6 +12,9 @@ import data.nat.order.lemmas /-! # Miscellaneous lemmas about the integers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains lemmas about integers, which require further imports than `data.int.basic` or `data.int.order`. diff --git a/src/data/int/modeq.lean b/src/data/int/modeq.lean index 9f63d07fa77e9..8c2d99cf0c206 100644 --- a/src/data/int/modeq.lean +++ b/src/data/int/modeq.lean @@ -10,6 +10,9 @@ import tactic.ring # Congruences modulo an integer +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the equivalence relation `a ≡ b [ZMOD n]` on the integers, similarly to how `data.nat.modeq` defines them for the natural numbers. The notation is short for `n.modeq a b`, which is defined to be `a % n = b % n` for integers `a b n`. diff --git a/src/data/int/nat_prime.lean b/src/data/int/nat_prime.lean index 4579ccc4dea97..c49b3c2905724 100644 --- a/src/data/int/nat_prime.lean +++ b/src/data/int/nat_prime.lean @@ -7,6 +7,9 @@ import data.nat.prime /-! # Lemmas about nat.prime using `int`s + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ open nat diff --git a/src/data/int/order/lemmas.lean b/src/data/int/order/lemmas.lean index 658dd4abab57b..9a3e5142a6891 100644 --- a/src/data/int/order/lemmas.lean +++ b/src/data/int/order/lemmas.lean @@ -9,6 +9,9 @@ import algebra.order.ring.abs /-! # Further lemmas about the integers + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. The distinction between this file and `data.int.order.basic` is not particularly clear. They are separated by now to minimize the porting requirements for tactics during the transition to mathlib4. After `data.rat.order` has been ported, please feel free to reorganize these two files. diff --git a/src/data/int/order/units.lean b/src/data/int/order/units.lean index 838e7fc8f63f1..d9970503fa589 100644 --- a/src/data/int/order/units.lean +++ b/src/data/int/order/units.lean @@ -10,6 +10,9 @@ import algebra.group_power.order /-! # Lemmas about units in `ℤ`, which interact with the order structure. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ namespace int diff --git a/src/data/int/sqrt.lean b/src/data/int/sqrt.lean index 829aa955bb9d2..72aaef66385da 100644 --- a/src/data/int/sqrt.lean +++ b/src/data/int/sqrt.lean @@ -8,6 +8,9 @@ import data.nat.sqrt /-! # Square root of integers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the square root function on integers. `int.sqrt z` is the greatest integer `r` such that `r * r ≤ z`. If `z ≤ 0`, then `int.sqrt z = 0`. -/ diff --git a/src/data/list/func.lean b/src/data/list/func.lean index 96224e56d8013..f075c0c84e28f 100644 --- a/src/data/list/func.lean +++ b/src/data/list/func.lean @@ -8,6 +8,9 @@ import data.nat.order.basic /-! # Lists as Functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Definitions for using lists as finite representations of finitely-supported functions with domain ℕ. diff --git a/src/data/nat/bits.lean b/src/data/nat/bits.lean index 984ffb50a06f7..662a5a507bef0 100644 --- a/src/data/nat/bits.lean +++ b/src/data/nat/bits.lean @@ -8,6 +8,9 @@ import data.nat.basic /-! # Additional properties of binary recursion on `nat` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file documents additional properties of binary recursion, which allows us to more easily work with operations which do depend on the number of leading zeros in the binary representation of `n`. diff --git a/src/data/nat/cast/field.lean b/src/data/nat/cast/field.lean index 8a2bc498cf4df..8b316e9f7e19c 100644 --- a/src/data/nat/cast/field.lean +++ b/src/data/nat/cast/field.lean @@ -10,6 +10,9 @@ import data.nat.cast.basic /-! # Cast of naturals into fields +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file concerns the canonical homomorphism `ℕ → F`, where `F` is a field. ## Main results diff --git a/src/data/nat/choose/basic.lean b/src/data/nat/choose/basic.lean index fc9d2fcc7e983..51c05e27848b5 100644 --- a/src/data/nat/choose/basic.lean +++ b/src/data/nat/choose/basic.lean @@ -8,6 +8,9 @@ import data.nat.factorial.basic /-! # Binomial coefficients +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines binomial coefficients and proves simple lemmas (i.e. those not requiring more imports). diff --git a/src/data/nat/choose/bounds.lean b/src/data/nat/choose/bounds.lean index f8102e4653ad0..3d9e5a32dd778 100644 --- a/src/data/nat/choose/bounds.lean +++ b/src/data/nat/choose/bounds.lean @@ -11,6 +11,9 @@ import data.nat.choose.basic /-! # Inequalities for binomial coefficients +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves exponential bounds on binomial coefficients. We might want to add here the bounds `n^r/r^r ≤ n.choose r ≤ e^r n^r/r^r` in the future. diff --git a/src/data/nat/choose/dvd.lean b/src/data/nat/choose/dvd.lean index 61dc0c6023377..11c3114090e97 100644 --- a/src/data/nat/choose/dvd.lean +++ b/src/data/nat/choose/dvd.lean @@ -8,6 +8,9 @@ import data.nat.prime /-! # Divisibility properties of binomial coefficients + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ namespace nat diff --git a/src/data/nat/dist.lean b/src/data/nat/dist.lean index 256b6f41b482c..5b32346de16e6 100644 --- a/src/data/nat/dist.lean +++ b/src/data/nat/dist.lean @@ -8,6 +8,9 @@ import data.nat.order.basic /-! # Distance function on ℕ +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines a simple distance function on naturals from truncated substraction. -/ diff --git a/src/data/nat/even_odd_rec.lean b/src/data/nat/even_odd_rec.lean index 4eb5627bf7d33..cb861068866c7 100644 --- a/src/data/nat/even_odd_rec.lean +++ b/src/data/nat/even_odd_rec.lean @@ -6,7 +6,10 @@ Authors: Stuart Presnell import data.nat.basic -/-! # A recursion principle based on even and odd numbers. -/ +/-! # A recursion principle based on even and odd numbers. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4.-/ namespace nat diff --git a/src/data/nat/factorial/basic.lean b/src/data/nat/factorial/basic.lean index d9fe35f593f04..ba3c8d0b388f9 100644 --- a/src/data/nat/factorial/basic.lean +++ b/src/data/nat/factorial/basic.lean @@ -9,6 +9,9 @@ import data.nat.pow /-! # Factorial and variants +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the factorial, along with the ascending and descending variants. ## Main declarations diff --git a/src/data/nat/log.lean b/src/data/nat/log.lean index b0b69c3f9e0c9..a2e46d0a572de 100644 --- a/src/data/nat/log.lean +++ b/src/data/nat/log.lean @@ -9,6 +9,9 @@ import tactic.by_contra /-! # Natural number logarithms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines two `ℕ`-valued analogs of the logarithm of `n` with base `b`: * `log b n`: Lower logarithm, or floor **log**. Greatest `k` such that `b^k ≤ n`. * `clog b n`: Upper logarithm, or **c**eil **log**. Least `k` such that `n ≤ b^k`. diff --git a/src/data/nat/modeq.lean b/src/data/nat/modeq.lean index 1f1dfc5bf2e15..d0510320fe193 100644 --- a/src/data/nat/modeq.lean +++ b/src/data/nat/modeq.lean @@ -9,6 +9,9 @@ import tactic.abel /-! # Congruences modulo a natural number +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the equivalence relation `a ≡ b [MOD n]` on the natural numbers, and proves basic properties about it such as the Chinese Remainder Theorem `modeq_and_modeq_iff_modeq_mul`. diff --git a/src/data/nat/pairing.lean b/src/data/nat/pairing.lean index 6f38130882395..0dad86e65afa3 100644 --- a/src/data/nat/pairing.lean +++ b/src/data/nat/pairing.lean @@ -11,6 +11,9 @@ import algebra.order.monoid.min_max /-! # Naturals pairing function +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines a pairing function for the naturals as follows: ```text 0 1 4 9 16 diff --git a/src/data/nat/pow.lean b/src/data/nat/pow.lean index d6392ebc24aff..219b088bccaf2 100644 --- a/src/data/nat/pow.lean +++ b/src/data/nat/pow.lean @@ -7,6 +7,9 @@ import algebra.group_power.order /-! # `nat.pow` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Results on the power operation on natural numbers. -/ diff --git a/src/data/nat/prime.lean b/src/data/nat/prime.lean index 261e2de73a035..57775b325ca8a 100644 --- a/src/data/nat/prime.lean +++ b/src/data/nat/prime.lean @@ -17,6 +17,9 @@ import tactic.by_contra /-! # Prime numbers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file deals with prime numbers: natural numbers `p ≥ 2` whose only divisors are `p` and `1`. ## Important declarations diff --git a/src/data/nat/size.lean b/src/data/nat/size.lean index ce81b1330f031..56ba762a64894 100644 --- a/src/data/nat/size.lean +++ b/src/data/nat/size.lean @@ -6,7 +6,10 @@ Authors: Floris van Doorn, Leonardo de Moura, Jeremy Avigad, Mario Carneiro import data.nat.pow import data.nat.bits -/-! Lemmas about `size`. -/ +/-! +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Lemmas about `size`. -/ namespace nat diff --git a/src/data/nat/sqrt.lean b/src/data/nat/sqrt.lean index e614c046c6ab0..d9ec06da13855 100644 --- a/src/data/nat/sqrt.lean +++ b/src/data/nat/sqrt.lean @@ -9,6 +9,9 @@ import data.nat.size /-! # Square root of natural numbers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines an efficient binary implementation of the square root function that returns the unique `r` such that `r * r ≤ n < (r + 1) * (r + 1)`. It takes advantage of the binary representation by replacing the multiplication by 2 appearing in diff --git a/src/data/nat/with_bot.lean b/src/data/nat/with_bot.lean index 6f1b74c9e3dae..713a12e2cc1f8 100644 --- a/src/data/nat/with_bot.lean +++ b/src/data/nat/with_bot.lean @@ -9,6 +9,9 @@ import algebra.order.monoid.with_top /-! # `with_bot ℕ` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Lemmas about the type of natural numbers with a bottom element adjoined. -/ diff --git a/src/data/part.lean b/src/data/part.lean index c85a0e1a1de88..a771023f74a83 100644 --- a/src/data/part.lean +++ b/src/data/part.lean @@ -9,6 +9,9 @@ import logic.equiv.defs /-! # Partial values of a type +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines `part α`, the partial values of a type. `o : part α` carries a proposition `o.dom`, its domain, along with a function `get : o.dom → α`, its diff --git a/src/data/pi/lex.lean b/src/data/pi/lex.lean index 847e495469617..6e7760e4c1e5d 100644 --- a/src/data/pi/lex.lean +++ b/src/data/pi/lex.lean @@ -11,6 +11,9 @@ import algebra.order.group.defs /-! # Lexicographic order on Pi types +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the lexicographic order for Pi types. `a` is less than `b` if `a i = b i` for all `i` up to some point `k`, and `a k < b k`. diff --git a/src/data/pnat/basic.lean b/src/data/pnat/basic.lean index 6afff1bc11ed0..95208af3aadf8 100644 --- a/src/data/pnat/basic.lean +++ b/src/data/pnat/basic.lean @@ -13,6 +13,9 @@ import algebra.order.positive.ring /-! # The positive natural numbers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file develops the type `ℕ+` or `pnat`, the subtype of natural numbers that are positive. It is defined in `data.pnat.defs`, but most of the development is deferred to here so that `data.pnat.defs` can have very few imports. diff --git a/src/data/pnat/find.lean b/src/data/pnat/find.lean index a1828db9d89eb..80ba4eafd2f56 100644 --- a/src/data/pnat/find.lean +++ b/src/data/pnat/find.lean @@ -8,6 +8,9 @@ import data.pnat.basic /-! # Explicit least witnesses to existentials on positive natural numbers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Implemented via calling out to `nat.find`. -/ diff --git a/src/data/rat/basic.lean b/src/data/rat/basic.lean index 1ad0a5bdba2ee..c251ff288ed30 100644 --- a/src/data/rat/basic.lean +++ b/src/data/rat/basic.lean @@ -9,6 +9,9 @@ import data.rat.defs /-! # Field Structure on the Rational Numbers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Summary We put the (discrete) field structure on the type `ℚ` of rational numbers that diff --git a/src/data/rat/defs.lean b/src/data/rat/defs.lean index 0419d00033201..e0b4580e5ecca 100644 --- a/src/data/rat/defs.lean +++ b/src/data/rat/defs.lean @@ -13,6 +13,9 @@ import data.pnat.defs /-! # Basics for the Rational Numbers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Summary We define the integral domain structure on `ℚ` and prove basic lemmas about it. diff --git a/src/data/rat/lemmas.lean b/src/data/rat/lemmas.lean index dca8988238a72..ed1458d007002 100644 --- a/src/data/rat/lemmas.lean +++ b/src/data/rat/lemmas.lean @@ -12,6 +12,9 @@ import tactic.nth_rewrite /-! # Further lemmas for the Rational Numbers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + -/ namespace rat diff --git a/src/data/rat/order.lean b/src/data/rat/order.lean index 46004118b583a..5024f984112f9 100644 --- a/src/data/rat/order.lean +++ b/src/data/rat/order.lean @@ -11,6 +11,9 @@ import tactic.assert_exists /-! # Order for Rational Numbers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Summary We define the order on `ℚ`, prove that `ℚ` is a discrete, linearly ordered field, and define diff --git a/src/data/rat/sqrt.lean b/src/data/rat/sqrt.lean index 40aa86f330016..82702e4ca6c03 100644 --- a/src/data/rat/sqrt.lean +++ b/src/data/rat/sqrt.lean @@ -10,6 +10,9 @@ import data.int.sqrt /-! # Square root on rational numbers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the square root function on rational numbers `rat.sqrt` and proves several theorems about it. diff --git a/src/data/rel.lean b/src/data/rel.lean index d3a782870872c..1c1b0f2186a25 100644 --- a/src/data/rel.lean +++ b/src/data/rel.lean @@ -9,6 +9,9 @@ import order.galois_connection /-! # Relations +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines bundled relations. A relation between `α` and `β` is a function `α → β → Prop`. Relations are also known as set-valued functions, or partial multifunctions. diff --git a/src/data/semiquot.lean b/src/data/semiquot.lean index 445adc1ef493c..50d0885522584 100644 --- a/src/data/semiquot.lean +++ b/src/data/semiquot.lean @@ -7,6 +7,9 @@ import data.set.lattice /-! # Semiquotients +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A data type for semiquotients, which are classically equivalent to nonempty sets, but are useful for programming; the idea is that a semiquotient set `S` represents some (particular but unknown) diff --git a/src/data/set/Union_lift.lean b/src/data/set/Union_lift.lean index c65156fc3296e..80b88373cbbed 100644 --- a/src/data/set/Union_lift.lean +++ b/src/data/set/Union_lift.lean @@ -7,6 +7,9 @@ import data.set.lattice import order.directed /-! # Union lift + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. This file defines `set.Union_lift` to glue together functions defined on each of a collection of sets to make a function on the Union of those sets. diff --git a/src/data/set/accumulate.lean b/src/data/set/accumulate.lean index 6e7eb323a9778..9e57338b44c66 100644 --- a/src/data/set/accumulate.lean +++ b/src/data/set/accumulate.lean @@ -7,6 +7,9 @@ import data.set.lattice /-! # Accumulate +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The function `accumulate` takes a set `s` and returns `⋃ y ≤ x, s y`. -/ diff --git a/src/data/set/enumerate.lean b/src/data/set/enumerate.lean index 9a47f76a99bda..8343cdeda7c1e 100644 --- a/src/data/set/enumerate.lean +++ b/src/data/set/enumerate.lean @@ -9,6 +9,9 @@ import data.nat.order.basic /-! # Set enumeration +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file allows enumeration of sets given a choice function. The definition does not assume `sel` actually is a choice function, i.e. `sel s ∈ s` and diff --git a/src/data/set/function.lean b/src/data/set/function.lean index 642d3f19b3a7b..a2e5964ac4191 100644 --- a/src/data/set/function.lean +++ b/src/data/set/function.lean @@ -9,6 +9,9 @@ import logic.function.conjugate /-! # Functions over sets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions ### Predicate diff --git a/src/data/set/functor.lean b/src/data/set/functor.lean index e39e448e1ddc5..a4ae66503ca4d 100644 --- a/src/data/set/functor.lean +++ b/src/data/set/functor.lean @@ -8,6 +8,9 @@ import data.set.lattice /-! # Functoriality of `set` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the functor structure of `set`. -/ diff --git a/src/data/set/intervals/basic.lean b/src/data/set/intervals/basic.lean index 1fa228c395e89..290b00e84d0fb 100644 --- a/src/data/set/intervals/basic.lean +++ b/src/data/set/intervals/basic.lean @@ -9,6 +9,9 @@ import data.set.prod /-! # Intervals +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In any preorder `α`, we define intervals (which on each side can be either infinite, open, or closed) using the following naming conventions: - `i`: infinite diff --git a/src/data/set/intervals/disjoint.lean b/src/data/set/intervals/disjoint.lean index f1109974c9763..c649e78d80052 100644 --- a/src/data/set/intervals/disjoint.lean +++ b/src/data/set/intervals/disjoint.lean @@ -8,6 +8,9 @@ import data.set.lattice /-! # Extra lemmas about intervals +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains lemmas about intervals that cannot be included into `data.set.intervals.basic` because this would create an `import` cycle. Namely, lemmas in this file can use definitions from `data.set.lattice`, including `disjoint`. diff --git a/src/data/set/intervals/group.lean b/src/data/set/intervals/group.lean index eb430b8c338aa..ad8805380b283 100644 --- a/src/data/set/intervals/group.lean +++ b/src/data/set/intervals/group.lean @@ -6,7 +6,10 @@ Authors: Johannes Hölzl, Mario Carneiro, Patrick Massot, Yury Kudryashov, Rémy import data.set.intervals.basic import algebra.order.group.abs -/-! ### Lemmas about arithmetic operations and intervals. -/ +/-! ### Lemmas about arithmetic operations and intervals. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4.-/ variables {α : Type*} diff --git a/src/data/set/intervals/monoid.lean b/src/data/set/intervals/monoid.lean index 27fba10575236..30d120a5f3a25 100644 --- a/src/data/set/intervals/monoid.lean +++ b/src/data/set/intervals/monoid.lean @@ -12,6 +12,9 @@ import algebra.group.basic /-! # Images of intervals under `(+ d)` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The lemmas in this file state that addition maps intervals bijectively. The typeclass `has_exists_add_of_le` is defined specifically to make them work when combined with `ordered_cancel_add_comm_monoid`; the lemmas below therefore apply to all diff --git a/src/data/set/intervals/ord_connected.lean b/src/data/set/intervals/ord_connected.lean index 72beaf11b0721..50952b855241b 100644 --- a/src/data/set/intervals/ord_connected.lean +++ b/src/data/set/intervals/ord_connected.lean @@ -9,6 +9,9 @@ import data.set.lattice /-! # Order-connected sets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We say that a set `s : set α` is `ord_connected` if for all `x y ∈ s` it includes the interval `[x, y]`. If `α` is a `densely_ordered` `conditionally_complete_linear_order` with the `order_topology`, then this condition is equivalent to `is_preconnected s`. If `α` is a diff --git a/src/data/set/intervals/order_iso.lean b/src/data/set/intervals/order_iso.lean index 5aa3d325ec02b..b477ced88c324 100644 --- a/src/data/set/intervals/order_iso.lean +++ b/src/data/set/intervals/order_iso.lean @@ -8,6 +8,9 @@ import order.hom.set /-! # Lemmas about images of intervals under order isomorphisms. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ variables {α β : Type*} diff --git a/src/data/set/intervals/pi.lean b/src/data/set/intervals/pi.lean index b9a28ce41fd23..1ce85fa64b8e2 100644 --- a/src/data/set/intervals/pi.lean +++ b/src/data/set/intervals/pi.lean @@ -9,6 +9,9 @@ import data.set.lattice /-! # Intervals in `pi`-space +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this we prove various simple lemmas about intervals in `Π i, α i`. Closed intervals (`Ici x`, `Iic x`, `Icc x y`) are equal to products of their projections to `α i`, while (semi-)open intervals usually include the corresponding products as proper subsets. diff --git a/src/data/set/intervals/proj_Icc.lean b/src/data/set/intervals/proj_Icc.lean index 3b6d07a28d574..871779f42b151 100644 --- a/src/data/set/intervals/proj_Icc.lean +++ b/src/data/set/intervals/proj_Icc.lean @@ -9,6 +9,9 @@ import data.set.intervals.basic /-! # Projection of a line onto a closed interval +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given a linearly ordered type `α`, in this file we define * `set.proj_Icc (a b : α) (h : a ≤ b)` to be the map `α → [a, b]` sending `(-∞, a]` to `a`, `[b, ∞)` diff --git a/src/data/set/intervals/surj_on.lean b/src/data/set/intervals/surj_on.lean index 2a55228f23675..8089ce42be73f 100644 --- a/src/data/set/intervals/surj_on.lean +++ b/src/data/set/intervals/surj_on.lean @@ -9,6 +9,9 @@ import data.set.function /-! # Monotone surjective functions are surjective on intervals +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A monotone surjective function sends any interval in the domain onto the interval with corresponding endpoints in the range. This is expressed in this file using `set.surj_on`, and provided for all permutations of interval endpoints. diff --git a/src/data/set/intervals/unordered_interval.lean b/src/data/set/intervals/unordered_interval.lean index f5fc402a3a656..7dd4efbc72cc0 100644 --- a/src/data/set/intervals/unordered_interval.lean +++ b/src/data/set/intervals/unordered_interval.lean @@ -9,6 +9,9 @@ import data.set.intervals.basic /-! # Intervals without endpoints ordering +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In any lattice `α`, we define `interval a b` to be `Icc (a ⊓ b) (a ⊔ b)`, which in a linear order is the set of elements lying between `a` and `b`. diff --git a/src/data/set/intervals/with_bot_top.lean b/src/data/set/intervals/with_bot_top.lean index 6f4c2ee095456..872f690fdba60 100644 --- a/src/data/set/intervals/with_bot_top.lean +++ b/src/data/set/intervals/with_bot_top.lean @@ -8,6 +8,9 @@ import data.set.intervals.basic /-! # Intervals in `with_top α` and `with_bot α` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove various lemmas about `set.image`s and `set.preimage`s of intervals under `coe : α → with_top α` and `coe : α → with_bot α`. -/ diff --git a/src/data/set/lattice.lean b/src/data/set/lattice.lean index 3cb4e4722162b..9f70b6e4d79a0 100644 --- a/src/data/set/lattice.lean +++ b/src/data/set/lattice.lean @@ -10,6 +10,9 @@ import order.galois_connection /-! # The set lattice +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file provides usual set notation for unions and intersections, a `complete_lattice` instance for `set α`, and some more set constructions. diff --git a/src/data/set/pairwise.lean b/src/data/set/pairwise.lean index ba87b6e47e941..510c54cb4e6cc 100644 --- a/src/data/set/pairwise.lean +++ b/src/data/set/pairwise.lean @@ -10,6 +10,9 @@ import data.set.lattice /-! # Relations holding pairwise +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file develops pairwise relations and defines pairwise disjoint indexed sets. We also prove many basic facts about `pairwise`. It is possible that an intermediate file, diff --git a/src/data/setoid/basic.lean b/src/data/setoid/basic.lean index f8e8c56b9a906..4e69c1aae666e 100644 --- a/src/data/setoid/basic.lean +++ b/src/data/setoid/basic.lean @@ -9,6 +9,9 @@ import order.galois_connection /-! # Equivalence relations +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the complete lattice of equivalence relations on a type, results about the inductively defined equivalence closure of a binary relation, and the analogues of some isomorphism theorems for quotients of arbitrary types. diff --git a/src/data/stream/init.lean b/src/data/stream/init.lean index 380462917deae..30c26cbdcd95b 100644 --- a/src/data/stream/init.lean +++ b/src/data/stream/init.lean @@ -10,6 +10,9 @@ import logic.function.basic /-! # Streams a.k.a. infinite lists a.k.a. infinite sequences +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file used to be in the core library. It was moved to `mathlib` and renamed to `init` to avoid name clashes. -/ diff --git a/src/dynamics/fixed_points/basic.lean b/src/dynamics/fixed_points/basic.lean index 8ab7ed997ebbb..2520af1f44024 100644 --- a/src/dynamics/fixed_points/basic.lean +++ b/src/dynamics/fixed_points/basic.lean @@ -10,6 +10,9 @@ import group_theory.perm.basic /-! # Fixed points of a self-map +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define * the predicate `is_fixed_pt f x := f x = x`; diff --git a/src/group_theory/group_action/group.lean b/src/group_theory/group_action/group.lean index 8747e9d145e65..02971812f5784 100644 --- a/src/group_theory/group_action/group.lean +++ b/src/group_theory/group_action/group.lean @@ -9,6 +9,9 @@ import group_theory.group_action.units /-! # Group actions applied to various types of group +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains lemmas about `smul` on `group_with_zero`, and `group`. -/ diff --git a/src/group_theory/group_action/opposite.lean b/src/group_theory/group_action/opposite.lean index c0a98a8b8185d..e66543bb1f8fc 100644 --- a/src/group_theory/group_action/opposite.lean +++ b/src/group_theory/group_action/opposite.lean @@ -9,6 +9,9 @@ import group_theory.group_action.defs /-! # Scalar actions on and by `Mᵐᵒᵖ` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the actions on the opposite type `has_smul R Mᵐᵒᵖ`, and actions by the opposite type, `has_smul Rᵐᵒᵖ M`. diff --git a/src/group_theory/group_action/pi.lean b/src/group_theory/group_action/pi.lean index eba8a1f1549e7..f0cf07e8d94d5 100644 --- a/src/group_theory/group_action/pi.lean +++ b/src/group_theory/group_action/pi.lean @@ -9,6 +9,9 @@ import group_theory.group_action.defs /-! # Pi instances for multiplicative actions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines instances for mul_action and related structures on Pi types. ## See also diff --git a/src/group_theory/group_action/prod.lean b/src/group_theory/group_action/prod.lean index 728f6886acc79..d373216994e15 100644 --- a/src/group_theory/group_action/prod.lean +++ b/src/group_theory/group_action/prod.lean @@ -9,6 +9,9 @@ import group_theory.group_action.defs /-! # Prod instances for additive and multiplicative actions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines instances for binary product of additive and multiplicative actions and provides scalar multiplication as a homomorphism from `α × β` to `β`. diff --git a/src/group_theory/perm/basic.lean b/src/group_theory/perm/basic.lean index d3e93cc4fb647..67dae3c1b4efd 100644 --- a/src/group_theory/perm/basic.lean +++ b/src/group_theory/perm/basic.lean @@ -11,6 +11,9 @@ import logic.function.iterate /-! # The group of permutations (self-equivalences) of a type `α` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the `group` structure on `equiv.perm α`. -/ universes u v diff --git a/src/group_theory/perm/via_embedding.lean b/src/group_theory/perm/via_embedding.lean index 432d161f65be8..d1b4dbc15c0e6 100644 --- a/src/group_theory/perm/via_embedding.lean +++ b/src/group_theory/perm/via_embedding.lean @@ -8,6 +8,9 @@ import logic.equiv.set /-! # `equiv.perm.via_embedding`, a noncomputable analogue of `equiv.perm.via_fintype_embedding`. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ variables {α β : Type*} diff --git a/src/group_theory/submonoid/basic.lean b/src/group_theory/submonoid/basic.lean index 8fa051273cbc5..a503eca7573f6 100644 --- a/src/group_theory/submonoid/basic.lean +++ b/src/group_theory/submonoid/basic.lean @@ -11,6 +11,9 @@ import group_theory.subsemigroup.basic /-! # Submonoids: definition and `complete_lattice` structure +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines bundled multiplicative and additive submonoids. We also define a `complete_lattice` structure on `submonoid`s, define the closure of a set as the minimal submonoid that includes this set, and prove a few results about extending properties from a dense set (i.e. diff --git a/src/group_theory/subsemigroup/basic.lean b/src/group_theory/subsemigroup/basic.lean index 28d6e513792c9..daefae477541c 100644 --- a/src/group_theory/subsemigroup/basic.lean +++ b/src/group_theory/subsemigroup/basic.lean @@ -11,6 +11,9 @@ import data.set_like.basic /-! # Subsemigroups: definition and `complete_lattice` structure +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines bundled multiplicative and additive subsemigroups. We also define a `complete_lattice` structure on `subsemigroup`s, and define the closure of a set as the minimal subsemigroup that includes this set. diff --git a/src/group_theory/subsemigroup/center.lean b/src/group_theory/subsemigroup/center.lean index 0f393774d6893..0ab9e2901c059 100644 --- a/src/group_theory/subsemigroup/center.lean +++ b/src/group_theory/subsemigroup/center.lean @@ -9,6 +9,9 @@ import group_theory.subsemigroup.operations /-! # Centers of magmas and semigroups +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions * `set.center`: the center of a magma diff --git a/src/group_theory/subsemigroup/centralizer.lean b/src/group_theory/subsemigroup/centralizer.lean index 21398b2b81421..8eadf8a074db3 100644 --- a/src/group_theory/subsemigroup/centralizer.lean +++ b/src/group_theory/subsemigroup/centralizer.lean @@ -9,6 +9,9 @@ import algebra.group_with_zero.units.lemmas /-! # Centralizers of magmas and semigroups +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions * `set.centralizer`: the centralizer of a subset of a magma diff --git a/src/group_theory/subsemigroup/operations.lean b/src/group_theory/subsemigroup/operations.lean index b856172503b49..3046644562618 100644 --- a/src/group_theory/subsemigroup/operations.lean +++ b/src/group_theory/subsemigroup/operations.lean @@ -11,6 +11,9 @@ import algebra.group.type_tags /-! # Operations on `subsemigroup`s +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define various operations on `subsemigroup`s and `mul_hom`s. ## Main definitions diff --git a/src/logic/equiv/nat.lean b/src/logic/equiv/nat.lean index 20d401352ca6f..ffe91f0fa86ee 100644 --- a/src/logic/equiv/nat.lean +++ b/src/logic/equiv/nat.lean @@ -8,6 +8,9 @@ import data.nat.pairing /-! # Equivalences involving `ℕ` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines some additional constructive equivalences using `encodable` and the pairing function on `ℕ`. -/ diff --git a/src/logic/equiv/set.lean b/src/logic/equiv/set.lean index f99ab279507a7..ce3c95ec44b43 100644 --- a/src/logic/equiv/set.lean +++ b/src/logic/equiv/set.lean @@ -9,6 +9,9 @@ import logic.equiv.defs /-! # Equivalences and sets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we provide lemmas linking equivalences to sets. Some notable definitions are: diff --git a/src/logic/small/basic.lean b/src/logic/small/basic.lean index 5fde86ccdf39f..605e4c9b07f6f 100644 --- a/src/logic/small/basic.lean +++ b/src/logic/small/basic.lean @@ -8,6 +8,9 @@ import logic.equiv.set /-! # Small types +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A type is `w`-small if there exists an equivalence to some `S : Type w`. We provide a noncomputable model `shrink α : Type w`, and `equiv_shrink α : α ≃ shrink α`. diff --git a/src/order/antichain.lean b/src/order/antichain.lean index e6713ffba2f9b..ab3d0e5e6bbc0 100644 --- a/src/order/antichain.lean +++ b/src/order/antichain.lean @@ -8,6 +8,9 @@ import data.set.pairwise /-! # Antichains +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines antichains. An antichain is a set where any two distinct elements are not related. If the relation is `(≤)`, this corresponds to incomparability and usual order antichains. If the relation is `G.adj` for `G : simple_graph α`, this corresponds to independent sets of `G`. diff --git a/src/order/atoms.lean b/src/order/atoms.lean index efea364ba2fc6..e7863c9974d32 100644 --- a/src/order/atoms.lean +++ b/src/order/atoms.lean @@ -9,6 +9,9 @@ import order.well_founded /-! # Atoms, Coatoms, and Simple Lattices +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This module defines atoms, which are minimal non-`⊥` elements in bounded lattices, simple lattices, which are lattices with only two elements, and related ideas. diff --git a/src/order/bounded.lean b/src/order/bounded.lean index b7e2634a12a81..052e94d418516 100644 --- a/src/order/bounded.lean +++ b/src/order/bounded.lean @@ -9,6 +9,9 @@ import data.set.intervals.basic /-! # Bounded and unbounded sets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We prove miscellaneous lemmas about bounded and unbounded sets. Many of these are just variations on the same ideas, or similar results with a few minor differences. The file is divided into these different general ideas. diff --git a/src/order/bounds/basic.lean b/src/order/bounds/basic.lean index 7f6acfdbd8296..e688e36d34499 100644 --- a/src/order/bounds/basic.lean +++ b/src/order/bounds/basic.lean @@ -10,6 +10,9 @@ import data.set.n_ary # Upper / lower bounds +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define: * `upper_bounds`, `lower_bounds` : the set of upper bounds (resp., lower bounds) of a set; diff --git a/src/order/bounds/order_iso.lean b/src/order/bounds/order_iso.lean index 0f6958c3b49bb..b90ea907458e5 100644 --- a/src/order/bounds/order_iso.lean +++ b/src/order/bounds/order_iso.lean @@ -8,6 +8,9 @@ import order.hom.set /-! # Order isomorhpisms and bounds. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ variables {α β : Type*} diff --git a/src/order/chain.lean b/src/order/chain.lean index 0102c14d1e85d..9c0087af10389 100644 --- a/src/order/chain.lean +++ b/src/order/chain.lean @@ -9,6 +9,9 @@ import data.set_like.basic /-! # Chains and flags +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines chains for an arbitrary relation and flags for an order and proves Hausdorff's Maximality Principle. diff --git a/src/order/closure.lean b/src/order/closure.lean index d0d83d61be378..02cbb145e3131 100644 --- a/src/order/closure.lean +++ b/src/order/closure.lean @@ -11,6 +11,9 @@ import order.hom.basic /-! # Closure operators between preorders +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define (bundled) closure operators on a preorder as monotone (increasing), extensive (inflationary) and idempotent functions. We define closed elements for the operator as elements which are fixed by it. diff --git a/src/order/complete_boolean_algebra.lean b/src/order/complete_boolean_algebra.lean index 01e1f01f04326..dd590f754bc04 100644 --- a/src/order/complete_boolean_algebra.lean +++ b/src/order/complete_boolean_algebra.lean @@ -10,6 +10,9 @@ import logic.equiv.set /-! # Frames, completely distributive lattices and Boolean algebras +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define and provide API for frames, completely distributive lattices and completely distributive Boolean algebras. diff --git a/src/order/complete_lattice.lean b/src/order/complete_lattice.lean index 4859186a088d1..a1cfade619ff4 100644 --- a/src/order/complete_lattice.lean +++ b/src/order/complete_lattice.lean @@ -12,6 +12,9 @@ import order.hom.basic /-! # Theory of complete lattices +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions * `Sup` and `Inf` are the supremum and the infimum of a set; diff --git a/src/order/complete_lattice_intervals.lean b/src/order/complete_lattice_intervals.lean index 5f4c5a5e0bf6b..370a5b8330f7c 100644 --- a/src/order/complete_lattice_intervals.lean +++ b/src/order/complete_lattice_intervals.lean @@ -8,6 +8,9 @@ import data.set.intervals.ord_connected /-! # Subtypes of conditionally complete linear orders +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we give conditions on a subset of a conditionally complete linear order, to ensure that the subtype is itself conditionally complete. diff --git a/src/order/conditionally_complete_lattice/basic.lean b/src/order/conditionally_complete_lattice/basic.lean index 9d6897de353ea..8eea89f3d5511 100644 --- a/src/order/conditionally_complete_lattice/basic.lean +++ b/src/order/conditionally_complete_lattice/basic.lean @@ -11,6 +11,9 @@ import data.set.lattice /-! # Theory of conditionally complete lattices. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A conditionally complete lattice is a lattice in which every non-empty bounded subset `s` has a least upper bound and a greatest lower bound, denoted below by `Sup s` and `Inf s`. Typical examples are `ℝ`, `ℕ`, and `ℤ` with their usual orders. diff --git a/src/order/conditionally_complete_lattice/group.lean b/src/order/conditionally_complete_lattice/group.lean index 2f95db0f42273..b9c5bcef8ad2c 100644 --- a/src/order/conditionally_complete_lattice/group.lean +++ b/src/order/conditionally_complete_lattice/group.lean @@ -9,6 +9,9 @@ import algebra.order.group.type_tags /-! # Conditionally complete lattices and groups. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + -/ diff --git a/src/order/copy.lean b/src/order/copy.lean index e33bf8c203dc6..77da359129d7b 100644 --- a/src/order/copy.lean +++ b/src/order/copy.lean @@ -8,6 +8,9 @@ import order.conditionally_complete_lattice.basic /-! # Tooling to make copies of lattice structures +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Sometimes it is useful to make a copy of a lattice structure where one replaces the data parts with provably equal definitions that have better definitional properties. diff --git a/src/order/cover.lean b/src/order/cover.lean index 5a9f6fa0729b4..e654312d37f98 100644 --- a/src/order/cover.lean +++ b/src/order/cover.lean @@ -9,6 +9,9 @@ import order.antisymmetrization /-! # The covering relation +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the covering relation in an order. `b` is said to cover `a` if `a < b` and there is no element in between. We say that `b` weakly covers `a` if `a ≤ b` and there is no element between `a` and `b`. In a partial order this is equivalent to `a ⋖ b ∨ a = b`, in a preorder this diff --git a/src/order/fixed_points.lean b/src/order/fixed_points.lean index 5ce5fa03dd506..99ae88801f222 100644 --- a/src/order/fixed_points.lean +++ b/src/order/fixed_points.lean @@ -9,6 +9,9 @@ import order.hom.order /-! # Fixed point construction on complete lattices +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file sets up the basic theory of fixed points of a monotone function in a complete lattice. ## Main definitions diff --git a/src/order/galois_connection.lean b/src/order/galois_connection.lean index d112f4b776746..31d946268a364 100644 --- a/src/order/galois_connection.lean +++ b/src/order/galois_connection.lean @@ -10,6 +10,9 @@ import order.hom.set /-! # Galois connections, insertions and coinsertions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Galois connections are order theoretic adjoints, i.e. a pair of functions `u` and `l`, such that `∀ a b, l a ≤ b ↔ a ≤ u b`. diff --git a/src/order/heyting/regular.lean b/src/order/heyting/regular.lean index fbacda636ccd0..321537b802e76 100644 --- a/src/order/heyting/regular.lean +++ b/src/order/heyting/regular.lean @@ -8,6 +8,9 @@ import order.galois_connection /-! # Heyting regular elements +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines Heyting regular elements, elements of an Heyting algebra that are their own double complement, and proves that they form a boolean algebra. diff --git a/src/order/hom/order.lean b/src/order/hom/order.lean index dd5786296ba29..29105fa08b152 100644 --- a/src/order/hom/order.lean +++ b/src/order/hom/order.lean @@ -10,6 +10,9 @@ import order.hom.basic /-! # Lattice structure on order homomorphisms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the lattice structure on order homomorphisms, which are bundled monotone functions. diff --git a/src/order/hom/set.lean b/src/order/hom/set.lean index aa77d38f09788..e2a7308b2f386 100644 --- a/src/order/hom/set.lean +++ b/src/order/hom/set.lean @@ -9,6 +9,9 @@ import data.set.image /-! # Order homomorphisms and sets + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ open order_dual diff --git a/src/order/initial_seg.lean b/src/order/initial_seg.lean index 8b0a9fdc8aa33..16841b0aeddbe 100644 --- a/src/order/initial_seg.lean +++ b/src/order/initial_seg.lean @@ -10,6 +10,9 @@ import order.well_founded /-! # Initial and principal segments +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines initial and principal segments. ## Main definitions diff --git a/src/order/lattice_intervals.lean b/src/order/lattice_intervals.lean index e26b862a81131..5897c15e515d8 100644 --- a/src/order/lattice_intervals.lean +++ b/src/order/lattice_intervals.lean @@ -9,6 +9,9 @@ import order.bounds.basic /-! # Intervals in Lattices +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we provide instances of lattice structures on intervals within lattices. Some of them depend on the order of the endpoints of the interval, and thus are not made global instances. These are probably not all of the lattice instances that could be placed on these diff --git a/src/order/modular_lattice.lean b/src/order/modular_lattice.lean index dedce4deaae03..96b99249a51e9 100644 --- a/src/order/modular_lattice.lean +++ b/src/order/modular_lattice.lean @@ -9,6 +9,9 @@ import order.lattice_intervals /-! # Modular Lattices +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines (semi)modular lattices, a kind of lattice useful in algebra. For examples, look to the subobject lattices of abelian groups, submodules, and ideals, or consider any distributive lattice. diff --git a/src/order/monotone/extension.lean b/src/order/monotone/extension.lean index ae27ab1e97a6b..4336bfbe6b939 100644 --- a/src/order/monotone/extension.lean +++ b/src/order/monotone/extension.lean @@ -8,6 +8,9 @@ import order.conditionally_complete_lattice.basic /-! # Extension of a monotone function from a set to the whole space +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove that if a function is monotone and is bounded on a set `s`, then it admits a monotone extension to the whole space. -/ diff --git a/src/order/monotone/odd.lean b/src/order/monotone/odd.lean index ca04a30e0a40e..54fa7633b1132 100644 --- a/src/order/monotone/odd.lean +++ b/src/order/monotone/odd.lean @@ -9,6 +9,9 @@ import algebra.order.group.instances /-! # Monotonicity of odd functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + An odd function on a linear ordered additive commutative group `G` is monotone on the whole group provided that is is monotone on `set.Ici 0`, see `monotone_of_odd_of_monotone_on_nonneg`. We also prove versions of this lemma for `antitone`, `strict_mono`, and `strict_anti`. diff --git a/src/order/monotone/union.lean b/src/order/monotone/union.lean index 56bb83e27f15e..f1fe9ac9d7aae 100644 --- a/src/order/monotone/union.lean +++ b/src/order/monotone/union.lean @@ -8,6 +8,9 @@ import order.bounds.basic /-! # Monotonicity on intervals +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove that a function is (strictly) monotone (or antitone) on a linear order `α` provided that it is (strictly) monotone on `(-∞, a]` and on `[a, +∞)`. This is a special case of a more general statement where one deduces monotonicity on a union from monotonicity on each diff --git a/src/order/ord_continuous.lean b/src/order/ord_continuous.lean index e334efbb4c920..d4a4f24ce34f7 100644 --- a/src/order/ord_continuous.lean +++ b/src/order/ord_continuous.lean @@ -9,6 +9,9 @@ import order.rel_iso.basic /-! # Order continuity +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We say that a function is *left order continuous* if it sends all least upper bounds to least upper bounds. The order dual notion is called *right order continuity*. diff --git a/src/order/semiconj_Sup.lean b/src/order/semiconj_Sup.lean index 9d2b2f621412d..d81a361900c44 100644 --- a/src/order/semiconj_Sup.lean +++ b/src/order/semiconj_Sup.lean @@ -13,6 +13,9 @@ import algebra.hom.equiv.units.basic /-! # Semiconjugate by `Sup` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove two facts about semiconjugate (families of) functions. First, if an order isomorphism `fa : α → α` is semiconjugate to an order embedding `fb : β → β` by diff --git a/src/order/succ_pred/basic.lean b/src/order/succ_pred/basic.lean index eff74e663ee47..4b540794fae56 100644 --- a/src/order/succ_pred/basic.lean +++ b/src/order/succ_pred/basic.lean @@ -10,6 +10,9 @@ import order.iterate /-! # Successor and predecessor +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines successor and predecessor orders. `succ a`, the successor of an element `a : α` is the least element greater than `a`. `pred a` is the greatest element less than `a`. Typical examples include `ℕ`, `ℤ`, `ℕ+`, `fin n`, but also `enat`, the lexicographic order of a successor/predecessor diff --git a/src/order/succ_pred/interval_succ.lean b/src/order/succ_pred/interval_succ.lean index c1711795ccd0c..79a635165d2a4 100644 --- a/src/order/succ_pred/interval_succ.lean +++ b/src/order/succ_pred/interval_succ.lean @@ -9,6 +9,9 @@ import order.succ_pred.basic /-! # Intervals `Ixx (f x) (f (order.succ x))` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove * `monotone.bUnion_Ico_Ioc_map_succ`: if `α` is a linear archimedean succ order and `β` is a linear diff --git a/src/order/succ_pred/relation.lean b/src/order/succ_pred/relation.lean index 15a145eeb1af0..31eacfbe6c891 100644 --- a/src/order/succ_pred/relation.lean +++ b/src/order/succ_pred/relation.lean @@ -7,6 +7,9 @@ import order.succ_pred.basic /-! # Relations on types with a `succ_order` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains properties about relations on types with a `succ_order` and their closure operations (like the transitive closure). -/ diff --git a/src/order/zorn.lean b/src/order/zorn.lean index 7318858d40674..7d1e922a3b667 100644 --- a/src/order/zorn.lean +++ b/src/order/zorn.lean @@ -8,6 +8,9 @@ import order.chain /-! # Zorn's lemmas +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves several formulations of Zorn's Lemma. ## Variants diff --git a/src/order/zorn_atoms.lean b/src/order/zorn_atoms.lean index f009f8d38cf93..e8f47f21475b7 100644 --- a/src/order/zorn_atoms.lean +++ b/src/order/zorn_atoms.lean @@ -9,6 +9,9 @@ import order.atoms /-! # Zorn lemma for (co)atoms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we use Zorn's lemma to prove that a partial order is atomic if every nonempty chain `c`, `⊥ ∉ c`, has a lower bound not equal to `⊥`. We also prove the order dual version of this statement. diff --git a/src/ring_theory/ore_localization/ore_set.lean b/src/ring_theory/ore_localization/ore_set.lean index ed9288e4dd111..5524547685e07 100644 --- a/src/ring_theory/ore_localization/ore_set.lean +++ b/src/ring_theory/ore_localization/ore_set.lean @@ -10,6 +10,9 @@ import group_theory.submonoid.basic # (Right) Ore sets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This defines right Ore sets on arbitrary monoids. ## References diff --git a/src/set_theory/cardinal/schroeder_bernstein.lean b/src/set_theory/cardinal/schroeder_bernstein.lean index a1dc7eee68bc7..411e0ec7a6247 100644 --- a/src/set_theory/cardinal/schroeder_bernstein.lean +++ b/src/set_theory/cardinal/schroeder_bernstein.lean @@ -9,6 +9,9 @@ import order.zorn /-! # Schröder-Bernstein theorem, well-ordering of cardinals +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves the Schröder-Bernstein theorem (see `schroeder_bernstein`), the well-ordering of cardinals (see `min_injective`) and the totality of their order (see `total`). From ab3a26d502fcc8f48ade690f3f1283cdd3136117 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 3 Jan 2023 04:08:33 +0000 Subject: [PATCH 0134/1291] chore(algebra/big_operators/basic): generalize `finset.prod_of_empty` (#18045) This also generalizes `prod_empty` to arbitrary `fintype` instances. --- src/algebra/big_operators/basic.lean | 9 +++++---- 1 file changed, 5 insertions(+), 4 deletions(-) diff --git a/src/algebra/big_operators/basic.lean b/src/algebra/big_operators/basic.lean index b8641c0685d44..8a7c4bbb5cd97 100644 --- a/src/algebra/big_operators/basic.lean +++ b/src/algebra/big_operators/basic.lean @@ -200,7 +200,8 @@ section comm_monoid variables [comm_monoid β] @[simp, to_additive] lemma prod_empty : ∏ x in ∅, f x = 1 := rfl -@[to_additive] lemma prod_of_empty [is_empty α] : ∏ i, f i = 1 := by rw [univ_eq_empty, prod_empty] +@[to_additive] lemma prod_of_empty [is_empty α] (s : finset α) : ∏ i in s, f i = 1 := +by rw [eq_empty_of_is_empty s, prod_empty] @[simp, to_additive] lemma prod_cons (h : a ∉ s) : (∏ x in (cons a s h), f x) = f a * ∏ x in s, f x := @@ -1607,13 +1608,13 @@ prod_bijective e e.bijective f g h variables {f s} @[to_additive] -lemma prod_unique {α β : Type*} [comm_monoid β] [unique α] (f : α → β) : +lemma prod_unique {α β : Type*} [comm_monoid β] [unique α] [fintype α] (f : α → β) : (∏ x : α, f x) = f default := by rw [univ_unique, prod_singleton] -@[to_additive] lemma prod_empty {α β : Type*} [comm_monoid β] [is_empty α] (f : α → β) : +@[to_additive] lemma prod_empty {α β : Type*} [comm_monoid β] [is_empty α] [fintype α] (f : α → β) : (∏ x : α, f x) = 1 := -by rw [eq_empty_of_is_empty (univ : finset α), finset.prod_empty] +finset.prod_of_empty _ @[to_additive] lemma prod_subsingleton {α β : Type*} [comm_monoid β] [subsingleton α] [fintype α] (f : α → β) (a : α) : From 6cb77a8eaff0ddd100e87b1591c6d3ad319514ff Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9mi=20Bottinelli?= Date: Tue, 3 Jan 2023 06:58:48 +0000 Subject: [PATCH 0135/1291] feat(category_theory/mittag_leffler): Introduce the Mittag-Leffler condition (#17905) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit - [x] depends on: #18013 Co-authored-by: Rémi Bottinelli Co-authored-by: Junyan Xu --- src/category_theory/mittag_leffler.lean | 224 ++++++++++++++++++++++++ src/order/directed.lean | 12 ++ 2 files changed, 236 insertions(+) create mode 100644 src/category_theory/mittag_leffler.lean diff --git a/src/category_theory/mittag_leffler.lean b/src/category_theory/mittag_leffler.lean new file mode 100644 index 0000000000000..eeb7324fb9beb --- /dev/null +++ b/src/category_theory/mittag_leffler.lean @@ -0,0 +1,224 @@ +/- +Copyright (c) 2022 Rémi Bottinelli, Junyan Xu. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Rémi Bottinelli, Junyan Xu +-/ +import category_theory.filtered +import data.set.finite + +/-! +# The Mittag-Leffler condition + +This files defines the Mittag-Leffler condition for cofiltered systems and (TODO) other properties +of such systems and their sections. + +## Main definitions + +Given a functor `F : J ⥤ Type v`: + +* For `j : J`, `F.eventual_range j` is the intersections of all ranges of morphisms `F.map f` + where `f` has codomain `j`. +* `F.is_mittag_leffler` states that the functor `F` satisfies the Mittag-Leffler + condition: the ranges of morphisms `F.map f` (with `f` having codomain `j`) stabilize. +* If `J` is cofiltered `F.to_eventual_ranges` is the subfunctor of `F` obtained by restriction + to `F.eventual_range`. +* `F.to_preimages` restricts a functor to preimages of a given set in some `F.obj i`. If `J` is + cofiltered, then it is Mittag-Leffler if `F` is, see `is_mittag_leffler.to_preimages`. + +## Main statements + +* `is_mittag_leffler_of_exists_finite_range` shows that if `J` is cofiltered and for all `j`, + there exists some `i` and `f : i ⟶ j` such that the range of `F.map f` is finite, then + `F` is Mittag-Leffler. +* `to_eventual_ranges_surjective` shows that if `F` is Mittag-Leffler, then `F.to_eventual_ranges` + has all morphisms `F.map f` surjective. + +## Todo + +* Specialize to inverse systems and fintype systems. +* Prove [Stacks: Lemma 0597](https://stacks.math.columbia.edu/tag/0597) + +## References + +* [Stacks: Mittag-Leffler systems](https://stacks.math.columbia.edu/tag/0594) + +## Tags + +Mittag-Leffler, surjective, eventual range, inverse system, + +-/ + +universes u v + +namespace category_theory +namespace functor + +open is_cofiltered set functor_to_types + +variables {J : Type u} [category J] (F : J ⥤ Type v) {i j k : J} (s : set (F.obj i)) + +/-- +The eventual range of the functor `F : J ⥤ Type v` at index `j : J` is the intersection +of the ranges of all maps `F.map f` with `i : J` and `f : i ⟶ j`. +-/ +def eventual_range (j : J) := ⋂ i (f : i ⟶ j), range (F.map f) + +lemma mem_eventual_range_iff {x : F.obj j} : + x ∈ F.eventual_range j ↔ ∀ ⦃i⦄ (f : i ⟶ j), x ∈ range (F.map f) := mem_Inter₂ + +/-- +The functor `F : J ⥤ Type v` satisfies the Mittag-Leffler condition if for all `j : J`, +there exists some `i : J` and `f : i ⟶ j` such that for all `k : J` and `g : k ⟶ j`, the range +of `F.map f` is contained in that of `F.map g`; +in other words (see `is_mittag_leffler_iff_eventual_range`), the eventual range at `j` is attained +by some `f : i ⟶ j`. +-/ +def is_mittag_leffler : Prop := +∀ j : J, ∃ i (f : i ⟶ j), ∀ ⦃k⦄ (g : k ⟶ j), range (F.map f) ⊆ range (F.map g) + +lemma is_mittag_leffler_iff_eventual_range : F.is_mittag_leffler ↔ + ∀ j : J, ∃ i (f : i ⟶ j), F.eventual_range j = range (F.map f) := +forall_congr $ λ j, exists₂_congr $ λ i f, + ⟨λ h, (Inter₂_subset _ _).antisymm $ subset_Inter₂ h, λ h, h ▸ Inter₂_subset⟩ + +lemma is_mittag_leffler.subset_image_eventual_range (h : F.is_mittag_leffler) (f : j ⟶ i) : + F.eventual_range i ⊆ F.map f '' (F.eventual_range j) := +begin + obtain ⟨k, g, hg⟩ := F.is_mittag_leffler_iff_eventual_range.1 h j, + rw hg, intros x hx, + obtain ⟨x, rfl⟩ := F.mem_eventual_range_iff.1 hx (g ≫ f), + refine ⟨_, ⟨x, rfl⟩, by simpa only [F.map_comp]⟩, +end + +lemma eventual_range_eq_range_precomp (f : i ⟶ j) (g : j ⟶ k) + (h : F.eventual_range k = range (F.map g)) : + F.eventual_range k = range (F.map $ f ≫ g) := +begin + apply subset_antisymm, + { apply Inter₂_subset, }, + { rw [h, F.map_comp], apply range_comp_subset_range, } +end + +lemma is_mittag_leffler_of_surjective + (h : ∀ (i j : J) (f : i ⟶ j), (F.map f).surjective) : F.is_mittag_leffler := +λ j, ⟨j, 𝟙 j, λ k g, by rw [map_id, types_id, range_id, (h k j g).range_eq]⟩ + +/-- The subfunctor of `F` obtained by restricting to the preimages of a set `s ∈ F.obj i`. -/ +@[simps] def to_preimages : J ⥤ Type v := +{ obj := λ j, ⋂ f : j ⟶ i, F.map f ⁻¹' s, + map := λ j k g, maps_to.restrict (F.map g) _ _ $ λ x h, begin + rw [mem_Inter] at h ⊢, intro f, + rw [← mem_preimage, preimage_preimage], + convert h (g ≫ f), rw F.map_comp, refl, + end, + map_id' := λ j, by { simp_rw F.map_id, ext, refl }, + map_comp' := λ j k l f g, by { simp_rw F.map_comp, refl } } + +variable [is_cofiltered_or_empty J] + +lemma eventual_range_maps_to (f : j ⟶ i) : + (F.eventual_range j).maps_to (F.map f) (F.eventual_range i) := +λ x hx, begin + rw mem_eventual_range_iff at hx ⊢, + intros k f', + obtain ⟨l, g, g', he⟩ := cospan f f', + obtain ⟨x, rfl⟩ := hx g, + rw [← map_comp_apply, he, F.map_comp], + exact ⟨_, rfl⟩, +end + +lemma is_mittag_leffler.eq_image_eventual_range (h : F.is_mittag_leffler) (f : j ⟶ i) : + F.eventual_range i = F.map f '' (F.eventual_range j) := +(h.subset_image_eventual_range F f).antisymm $ maps_to'.1 (F.eventual_range_maps_to f) + +lemma eventual_range_eq_iff {f : i ⟶ j} : + F.eventual_range j = range (F.map f) ↔ + ∀ ⦃k⦄ (g : k ⟶ i), range (F.map f) ⊆ range (F.map $ g ≫ f) := +begin + rw [subset_antisymm_iff, eventual_range, and_iff_right (Inter₂_subset _ _), subset_Inter₂_iff], + refine ⟨λ h k g, h _ _, λ h j' f', _⟩, + obtain ⟨k, g, g', he⟩ := cospan f f', + refine (h g).trans _, + rw [he, F.map_comp], + apply range_comp_subset_range, +end + +lemma is_mittag_leffler_iff_subset_range_comp : F.is_mittag_leffler ↔ + ∀ j : J, ∃ i (f : i ⟶ j), ∀ ⦃k⦄ (g : k ⟶ i), range (F.map f) ⊆ range (F.map $ g ≫ f) := +by simp_rw [is_mittag_leffler_iff_eventual_range, eventual_range_eq_iff] + +lemma is_mittag_leffler.to_preimages (h : F.is_mittag_leffler) : + (F.to_preimages s).is_mittag_leffler := +(is_mittag_leffler_iff_subset_range_comp _).2 $ λ j, begin + obtain ⟨j₁, g₁, f₁, -⟩ := cone_objs i j, + obtain ⟨j₂, f₂, h₂⟩ := F.is_mittag_leffler_iff_eventual_range.1 h j₁, + refine ⟨j₂, f₂ ≫ f₁, λ j₃ f₃, _⟩, + rintro _ ⟨⟨x, hx⟩, rfl⟩, + have : F.map f₂ x ∈ F.eventual_range j₁, { rw h₂, exact ⟨_, rfl⟩ }, + obtain ⟨y, hy, h₃⟩ := h.subset_image_eventual_range F (f₃ ≫ f₂) this, + refine ⟨⟨y, mem_Inter.2 $ λ g₂, _⟩, subtype.ext _⟩, + { obtain ⟨j₄, f₄, h₄⟩ := cone_maps g₂ ((f₃ ≫ f₂) ≫ g₁), + obtain ⟨y, rfl⟩ := F.mem_eventual_range_iff.1 hy f₄, + rw ← map_comp_apply at h₃, + rw [mem_preimage, ← map_comp_apply, h₄, ← category.assoc, map_comp_apply, h₃, ← map_comp_apply], + apply mem_Inter.1 hx }, + { simp_rw [to_preimages_map, maps_to.coe_restrict_apply, subtype.coe_mk], + rw [← category.assoc, map_comp_apply, h₃, map_comp_apply] }, +end + +lemma is_mittag_leffler_of_exists_finite_range + (h : ∀ (j : J), ∃ i (f : i ⟶ j), (range $ F.map f).finite) : + F.is_mittag_leffler := +λ j, begin + obtain ⟨i, hi, hf⟩ := h j, + obtain ⟨m, ⟨i, f, hm⟩, hmin⟩ := finset.is_well_founded_lt.wf.has_min + {s : finset (F.obj j) | ∃ i (f : i ⟶ j), ↑s = range (F.map f)} ⟨_, i, hi, hf.coe_to_finset⟩, + refine ⟨i, f, λ k g, + (directed_on_range.mp $ F.ranges_directed j).is_bot_of_is_min ⟨⟨i, f⟩, rfl⟩ _ _ ⟨⟨k, g⟩, rfl⟩⟩, + rintro _ ⟨⟨k', g'⟩, rfl⟩ hl, + refine (eq_of_le_of_not_lt hl _).ge, + have := hmin _ ⟨k', g', (m.finite_to_set.subset $ hm.substr hl).coe_to_finset⟩, + rwa [finset.lt_iff_ssubset, ← finset.coe_ssubset, set.finite.coe_to_finset, hm] at this, +end + +/-- +The subfunctor of `F` obtained by restricting to the eventual range at each index. +-/ +@[simps] def to_eventual_ranges : J ⥤ Type v := +{ obj := λ j, F.eventual_range j, + map := λ i j f, (F.eventual_range_maps_to f).restrict _ _ _, + map_id' := λ i, by { simp_rw F.map_id, ext, refl }, + map_comp' := λ _ _ _ _ _, by { simp_rw F.map_comp, refl } } + +/-- +The sections of the functor `F : J ⥤ Type v` are in bijection with the sections of +`F.eventual_ranges`. +-/ +def to_eventual_ranges_sections_equiv : F.to_eventual_ranges.sections ≃ F.sections := +{ to_fun := λ s, ⟨_, λ i j f, subtype.coe_inj.2 $ s.prop f⟩, + inv_fun := λ s, ⟨λ j, ⟨_, mem_Inter₂.2 $ λ i f, ⟨_, s.prop f⟩⟩, λ i j f, subtype.ext $ s.prop f⟩, + left_inv := λ _, by { ext, refl }, + right_inv := λ _, by { ext, refl } } + +/-- +If `F` satisfies the Mittag-Leffler condition, its restriction to eventual ranges is a surjective +functor. +-/ +lemma surjective_to_eventual_ranges (h : F.is_mittag_leffler) (f : i ⟶ j) : + (F.to_eventual_ranges.map f).surjective := +λ ⟨x, hx⟩, by { obtain ⟨y, hy, rfl⟩ := h.subset_image_eventual_range F f hx, exact ⟨⟨y, hy⟩, rfl⟩ } + +/-- If `F` is nonempty at each index and Mittag-Leffler, then so is `F.to_eventual_ranges`. -/ +lemma to_eventual_ranges_nonempty (h : F.is_mittag_leffler) [∀ (j : J), nonempty (F.obj j)] + (j : J) : nonempty (F.to_eventual_ranges.obj j) := +let ⟨i, f, h⟩ := F.is_mittag_leffler_iff_eventual_range.1 h j in +by { rw [to_eventual_ranges_obj, h], apply_instance } + +/-- If `F` has all arrows surjective, then it "factors through a poset". -/ +lemma thin_diagram_of_surjective (Fsur : ∀ (i j : J) (f : i ⟶ j), (F.map f).surjective) + (i j) (f g : i ⟶ j) : F.map f = F.map g := +let ⟨k, φ, hφ⟩ := cone_maps f g in +(Fsur k i φ).injective_comp_right $ by simp_rw [← types_comp, ← F.map_comp, hφ] + +end functor +end category_theory diff --git a/src/order/directed.lean b/src/order/directed.lean index 1d84b95013f1f..2bb601c877ab8 100644 --- a/src/order/directed.lean +++ b/src/order/directed.lean @@ -46,6 +46,10 @@ by simp [directed, directed_on]; refine ball_congr (λ x hx, by simp; refl) alias directed_on_iff_directed ↔ directed_on.directed_coe _ +theorem directed_on_range {f : β → α} : + directed r f ↔ directed_on r (set.range f) := +by simp_rw [directed, directed_on, set.forall_range_iff, set.exists_range_iff] + theorem directed_on_image {s} {f : β → α} : directed_on r (f '' s) ↔ directed_on (f ⁻¹'o r) s := by simp only [directed_on, set.ball_image_iff, set.bex_image_iff, order.preimage] @@ -166,6 +170,14 @@ protected lemma is_min.is_bot [is_directed α (≥)] (h : is_min a) : is_bot a : protected lemma is_max.is_top [is_directed α (≤)] (h : is_max a) : is_top a := h.to_dual.is_bot +lemma directed_on.is_bot_of_is_min {s : set α} (hd : directed_on (≥) s) + {m} (hm : m ∈ s) (hmin : ∀ a ∈ s, a ≤ m → m ≤ a) : ∀ a ∈ s, m ≤ a := +λ a as, let ⟨x, xs, xm, xa⟩ := hd m hm a as in (hmin x xs xm).trans xa + +lemma directed_on.is_top_of_is_max {s : set α} (hd : directed_on (≤) s) + {m} (hm : m ∈ s) (hmax : ∀ a ∈ s, m ≤ a → a ≤ m) : ∀ a ∈ s, a ≤ m := +@directed_on.is_bot_of_is_min αᵒᵈ _ s hd m hm hmax + lemma is_top_or_exists_gt [is_directed α (≤)] (a : α) : is_top a ∨ (∃ b, a < b) := (em (is_max a)).imp is_max.is_top not_is_max_iff.mp From ce967b52b40b6f8a6c0ee9af4c4135993a773040 Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Tue, 3 Jan 2023 09:58:32 +0000 Subject: [PATCH 0136/1291] feat(analysis/bounded_variation): remove [nonempty] from comp_anti/mono lemmas + golf (#18037) --- src/analysis/bounded_variation.lean | 59 +++++++++++++---------------- 1 file changed, 26 insertions(+), 33 deletions(-) diff --git a/src/analysis/bounded_variation.lean b/src/analysis/bounded_variation.lean index 23a659043b5cc..f2368959ba082 100644 --- a/src/analysis/bounded_variation.lean +++ b/src/analysis/bounded_variation.lean @@ -89,13 +89,7 @@ end lemma sum_le (f : α → E) {s : set α} (n : ℕ) {u : ℕ → α} (hu : monotone u) (us : ∀ i, u i ∈ s) : ∑ i in finset.range n, edist (f (u (i+1))) (f (u i)) ≤ evariation_on f s := -begin - let p : ℕ × {u : ℕ → α // monotone u ∧ ∀ i, u i ∈ s} := (n, ⟨u, hu, us⟩), - change ∑ i in finset.range p.1, edist (f ((p.2 : ℕ → α) (i+1))) (f ((p.2 : ℕ → α) i)) - ≤ evariation_on f s, - exact le_supr (λ (p : ℕ × {u : ℕ → α // monotone u ∧ ∀ i, u i ∈ s}), - ∑ i in finset.range p.1, edist (f ((p.2 : ℕ → α) (i+1))) (f ((p.2 : ℕ → α) i))) _, -end +le_supr_of_le ⟨n, u, hu, us⟩ le_rfl lemma sum_le_of_monotone_on_Iic (f : α → E) {s : set α} {n : ℕ} {u : ℕ → α} (hu : monotone_on u (Iic n)) @@ -537,41 +531,31 @@ end lemma comp_le_of_monotone_on (f : α → E) {s : set α} {t : set β} (φ : β → α) (hφ : monotone_on φ t) (φst : set.maps_to φ t s) : evariation_on (f ∘ φ) t ≤ evariation_on f s := -begin - apply supr_le _, - rintro ⟨n, ⟨u, hu, ut⟩⟩, - exact le_supr (λ (p : ℕ × {u : ℕ → α // monotone u ∧ ∀ i, u i ∈ s}), - ∑ i in finset.range p.1, edist (f ((p.2 : ℕ → α) (i+1))) (f ((p.2 : ℕ → α) i))) - ⟨n, ⟨φ ∘ u, λ x y xy, hφ (ut x) (ut y) (hu xy), λ i, φst (ut i)⟩⟩, -end +supr_le $ λ ⟨n, u, hu, ut⟩, le_supr_of_le + ⟨n, φ ∘ u, λ x y xy, hφ (ut x) (ut y) (hu xy), λ i, φst (ut i)⟩ le_rfl lemma comp_le_of_antitone_on (f : α → E) {s : set α} {t : set β} (φ : β → α) (hφ : antitone_on φ t) (φst : set.maps_to φ t s) : evariation_on (f ∘ φ) t ≤ evariation_on f s := begin - apply supr_le _, - rintros ⟨n, ⟨u, hu, ut⟩⟩, - change ∑ i in finset.range n, edist (f ∘ φ $ u (i+1)) (f ∘ φ $ u i) ≤ evariation_on f s, + refine supr_le _, + rintro ⟨n, u, hu, ut⟩, rw ←finset.sum_range_reflect, - have : ∀ x : ℕ, x ∈ finset.range n → - edist ((f ∘ φ) (u (n - 1 - x + 1))) ((f ∘ φ) (u (n - 1 - x))) = - edist ((f ∘ φ) (u (n - (x + 1)))) ((f ∘ φ) (u (n - x))) := λ x hx, by - { rw [edist_comm, nat.sub_sub, add_comm, nat.sub_succ, nat.add_one, nat.succ_pred_eq_of_pos], - simpa only [tsub_pos_iff_lt, finset.mem_range] using hx, }, - rw finset.sum_congr rfl this, - let ru : ℕ → β := λ i, u (n-i), - have rut : ∀ i : ℕ, ru i ∈ t := λ i, ut (n-i), - have hru : antitone ru := λ i j l, hu (n.sub_le_sub_left l), - exact le_supr (λ (p : ℕ × {u : ℕ → α // monotone u ∧ ∀ i, u i ∈ s}), - ∑ i in finset.range p.1, edist (f ((p.2 : ℕ → α) (i+1))) (f ((p.2 : ℕ → α) i))) - ⟨n, ⟨φ ∘ ru, λ x y xy, hφ (rut y) (rut x) (hru xy), λ i, φst (rut i)⟩⟩, + refine (finset.sum_congr rfl $ λ x hx, _).trans_le (le_supr_of_le ⟨n, λ i, φ (u $ n-i), + λ x y xy, hφ (ut _) (ut _) (hu $ n.sub_le_sub_left xy), λ i, φst (ut _)⟩ le_rfl), + dsimp only [subtype.coe_mk], + rw [edist_comm, nat.sub_sub, add_comm, nat.sub_succ, nat.add_one, nat.succ_pred_eq_of_pos], + simpa only [tsub_pos_iff_lt, finset.mem_range] using hx, end -lemma comp_eq_of_monotone_on (f : α → E) {s : set α} {t : set β} [nonempty β] (φ : β → α) +lemma comp_eq_of_monotone_on (f : α → E) {s : set α} {t : set β} (φ : β → α) (hφ : monotone_on φ t) (φst : set.maps_to φ t s) (φsur : set.surj_on φ t s) : evariation_on (f ∘ φ) t = evariation_on f s := begin apply le_antisymm (comp_le_of_monotone_on f φ hφ φst), + casesI is_empty_or_nonempty β, + { convert zero_le _, + exact evariation_on.subsingleton f ((subsingleton_of_subsingleton.image _).anti φsur) }, let ψ := φ.inv_fun_on t, have ψφs : set.eq_on (φ ∘ ψ) id s := φsur.right_inv_on_inv_fun_on, have ψts : set.maps_to ψ s t := φsur.maps_to_inv_fun_on, @@ -579,14 +563,17 @@ begin function.monotone_on_of_right_inv_on_of_maps_to hφ ψφs ψts, change evariation_on (f ∘ id) s ≤ evariation_on (f ∘ φ) t, rw ←eq_of_eq_on (ψφs.comp_left : set.eq_on (f ∘ (φ ∘ ψ)) (f ∘ id) s), - apply comp_le_of_monotone_on _ ψ hψ ψts, + exact comp_le_of_monotone_on _ ψ hψ ψts, end -lemma comp_eq_of_antitone_on (f : α → E) {s : set α} {t : set β} [nonempty β] (φ : β → α) +lemma comp_eq_of_antitone_on (f : α → E) {s : set α} {t : set β} (φ : β → α) (hφ : antitone_on φ t) (φst : set.maps_to φ t s) (φsur : set.surj_on φ t s) : evariation_on (f ∘ φ) t = evariation_on f s := begin apply le_antisymm (comp_le_of_antitone_on f φ hφ φst), + casesI is_empty_or_nonempty β, + { convert zero_le _, + exact evariation_on.subsingleton f ((subsingleton_of_subsingleton.image _).anti φsur) }, let ψ := φ.inv_fun_on t, have ψφs : set.eq_on (φ ∘ ψ) id s := φsur.right_inv_on_inv_fun_on, have ψts : set.maps_to ψ s t := φsur.maps_to_inv_fun_on, @@ -594,9 +581,15 @@ begin function.antitone_on_of_right_inv_on_of_maps_to hφ ψφs ψts, change evariation_on (f ∘ id) s ≤ evariation_on (f ∘ φ) t, rw ←eq_of_eq_on (ψφs.comp_left : set.eq_on (f ∘ (φ ∘ ψ)) (f ∘ id) s), - apply comp_le_of_antitone_on _ ψ hψ ψts, + exact comp_le_of_antitone_on _ ψ hψ ψts, end +open order_dual + +lemma comp_of_dual (f : α → E) (s : set α) : + evariation_on (f ∘ of_dual) (of_dual ⁻¹' s) = evariation_on f s := +comp_eq_of_antitone_on f of_dual (λ _ _ _ _, id) (maps_to_preimage _ _) $ λ x hx, ⟨x, hx, rfl⟩ + end evariation_on /-! ## Monotone functions and bounded variation -/ From 44b58b42794e5abe2bf86397c38e26b587e07e59 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 3 Jan 2023 11:06:42 +0000 Subject: [PATCH 0137/1291] feat(topology/algebra/module/multilinear): add continuous versions of constructions (#18044) One of these is a much more general case of `continuous_multilinear_map.curry0`; but to keep this PR small we don't attempt to remove `curry0`, instead just golfing it. --- src/analysis/normed_space/multilinear.lean | 45 ++++++++++++++++---- src/linear_algebra/alternating.lean | 2 +- src/linear_algebra/multilinear/basic.lean | 4 +- src/linear_algebra/pi_tensor_product.lean | 2 +- src/topology/algebra/module/multilinear.lean | 20 +++++++++ 5 files changed, 60 insertions(+), 13 deletions(-) diff --git a/src/analysis/normed_space/multilinear.lean b/src/analysis/normed_space/multilinear.lean index aba5f7d1d9f1c..649142203cb35 100644 --- a/src/analysis/normed_space/multilinear.lean +++ b/src/analysis/normed_space/multilinear.lean @@ -439,6 +439,40 @@ begin convert norm_le_pi_norm (λ j, f j m) i } end +section +variables (𝕜 G) + +@[simp] lemma norm_of_subsingleton [subsingleton ι] [nontrivial G] (i' : ι) : + ‖of_subsingleton 𝕜 G i'‖ = 1 := +begin + apply le_antisymm, + { refine op_norm_le_bound _ zero_le_one (λ m, _), + rw [fintype.prod_subsingleton _ i', one_mul, of_subsingleton_apply] }, + { obtain ⟨g, hg⟩ := exists_ne (0 : G), + rw ←norm_ne_zero_iff at hg, + have := (of_subsingleton 𝕜 G i').ratio_le_op_norm (λ _, g), + rwa [fintype.prod_subsingleton _ i', of_subsingleton_apply, div_self hg] at this }, +end + +@[simp] lemma nnnorm_of_subsingleton [subsingleton ι] [nontrivial G] (i' : ι) : + ‖of_subsingleton 𝕜 G i'‖₊ = 1 := +nnreal.eq $ norm_of_subsingleton _ _ _ + +variables {G} (E) + +@[simp] lemma norm_const_of_is_empty [is_empty ι] (x : G) : ‖const_of_is_empty 𝕜 E x‖ = ‖x‖ := +begin + apply le_antisymm, + { refine op_norm_le_bound _ (norm_nonneg _) (λ x, _), + rw [fintype.prod_empty, mul_one, const_of_is_empty_apply], }, + { simpa using (const_of_is_empty 𝕜 E x).le_op_norm 0 } +end + +@[simp] lemma nnnorm_const_of_is_empty [is_empty ι] (x : G) : ‖const_of_is_empty 𝕜 E x‖₊ = ‖x‖₊ := +nnreal.eq $ norm_const_of_is_empty _ _ _ + +end + section variables (𝕜 E E' G G') @@ -1264,10 +1298,7 @@ variables (𝕜 G) /-- Associating to an element `x` of a vector space `E₂` the continuous multilinear map in `0` variables taking the (unique) value `x` -/ def continuous_multilinear_map.curry0 (x : G') : G [×0]→L[𝕜] G' := -{ to_fun := λm, x, - map_add' := λ m i, fin.elim0 i, - map_smul' := λ m i, fin.elim0 i, - cont := continuous_const } +continuous_multilinear_map.const_of_is_empty 𝕜 _ x variable {G} @[simp] lemma continuous_multilinear_map.curry0_apply (x : G') (m : (fin 0) → G) : @@ -1291,11 +1322,7 @@ variables (𝕜 G) @[simp] lemma continuous_multilinear_map.curry0_norm (x : G') : ‖continuous_multilinear_map.curry0 𝕜 G x‖ = ‖x‖ := -begin - apply le_antisymm, - { exact continuous_multilinear_map.op_norm_le_bound _ (norm_nonneg _) (λm, by simp) }, - { simpa using (continuous_multilinear_map.curry0 𝕜 G x).le_op_norm 0 } -end +norm_const_of_is_empty _ _ _ variables {𝕜 G} @[simp] lemma continuous_multilinear_map.fin0_apply_norm (f : G [×0]→L[𝕜] G') {x : fin 0 → G} : diff --git a/src/linear_algebra/alternating.lean b/src/linear_algebra/alternating.lean index ef26176f204d0..d720e6127d3e7 100644 --- a/src/linear_algebra/alternating.lean +++ b/src/linear_algebra/alternating.lean @@ -289,7 +289,7 @@ def of_subsingleton [subsingleton ι] (i : ι) : alternating_map R M M ι := def const_of_is_empty [is_empty ι] (m : N) : alternating_map R M N ι := { to_fun := function.const _ m, map_eq_zero_of_eq' := λ v, is_empty_elim, - ..multilinear_map.const_of_is_empty R m } + ..multilinear_map.const_of_is_empty R _ m } end diff --git a/src/linear_algebra/multilinear/basic.lean b/src/linear_algebra/multilinear/basic.lean index b683759e9b64f..f2df2f37429c2 100644 --- a/src/linear_algebra/multilinear/basic.lean +++ b/src/linear_algebra/multilinear/basic.lean @@ -213,7 +213,7 @@ def of_subsingleton [subsingleton ι] (i' : ι) : multilinear_map R (λ _ : ι, map_smul' := λ m i r x, by { rw subsingleton.elim i i', simp only [function.eval, function.update_same], } } -variables {M₂} +variables (M₁) {M₂} /-- The constant map is multilinear when `ι` is empty. -/ @[simps {fully_applied := ff}] @@ -777,7 +777,7 @@ def dom_dom_congr_linear_equiv' {ι' : Type*} [decidable_eq ι'] (σ : ι ≃ ι /-- The space of constant maps is equivalent to the space of maps that are multilinear with respect to an empty family. -/ @[simps] def const_linear_equiv_of_is_empty [is_empty ι] : M₂ ≃ₗ[R] multilinear_map R M₁ M₂ := -{ to_fun := multilinear_map.const_of_is_empty R, +{ to_fun := multilinear_map.const_of_is_empty R _, map_add' := λ x y, rfl, map_smul' := λ t x, rfl, inv_fun := λ f, f 0, diff --git a/src/linear_algebra/pi_tensor_product.lean b/src/linear_algebra/pi_tensor_product.lean index 451d7d9408fbd..d50bc7c91cf76 100644 --- a/src/linear_algebra/pi_tensor_product.lean +++ b/src/linear_algebra/pi_tensor_product.lean @@ -444,7 +444,7 @@ variables (ι) /-- The tensor product over an empty index type `ι` is isomorphic to the base ring. -/ @[simps symm_apply] def is_empty_equiv [is_empty ι] : ⨂[R] i : ι, M ≃ₗ[R] R := -{ to_fun := lift (const_of_is_empty R 1), +{ to_fun := lift (const_of_is_empty R _ 1), inv_fun := λ r, r • tprod R (@is_empty_elim _ _ _), left_inv := λ x, by { apply x.induction_on, diff --git a/src/topology/algebra/module/multilinear.lean b/src/topology/algebra/module/multilinear.lean index 8bdb55433e7a6..3d3976403f62b 100644 --- a/src/topology/algebra/module/multilinear.lean +++ b/src/topology/algebra/module/multilinear.lean @@ -208,6 +208,26 @@ lemma pi_apply {ι' : Type*} {M' : ι' → Type*} [Π i, add_comm_monoid (M' i)] pi f m j = f j m := rfl +section +variables (R M₂) + +/-- The evaluation map from `ι → M₂` to `M₂` is multilinear at a given `i` when `ι` is subsingleton. +-/ +@[simps to_multilinear_map apply] +def of_subsingleton [subsingleton ι] (i' : ι) : continuous_multilinear_map R (λ _ : ι, M₂) M₂ := +{ to_multilinear_map := multilinear_map.of_subsingleton R _ i', + cont := continuous_apply _ } + +variables (M₁) {M₂} + +/-- The constant map is multilinear when `ι` is empty. -/ +@[simps to_multilinear_map apply] +def const_of_is_empty [is_empty ι] (m : M₂) : continuous_multilinear_map R M₁ M₂ := +{ to_multilinear_map := multilinear_map.const_of_is_empty R _ m, + cont := continuous_const } + +end + /-- If `g` is continuous multilinear and `f` is a collection of continuous linear maps, then `g (f₁ m₁, ..., fₙ mₙ)` is again a continuous multilinear map, that we call `g.comp_continuous_linear_map f`. -/ From 3387923e23bb1abbef4b4169eadfa0702770ac59 Mon Sep 17 00:00:00 2001 From: Patrick Massot Date: Tue, 3 Jan 2023 21:27:54 +0000 Subject: [PATCH 0138/1291] chore(data/fin/basic): Remove elab_strategy (#18049) --- src/data/fin/basic.lean | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/src/data/fin/basic.lean b/src/data/fin/basic.lean index bce40a81832ad..2fe63657d1da1 100644 --- a/src/data/fin/basic.lean +++ b/src/data/fin/basic.lean @@ -1237,7 +1237,7 @@ end /-- Define `f : Π i : fin n.succ, C i` by separately handling the cases `i = fin.last n` and `i = j.cast_succ`, `j : fin n`. -/ -@[elab_as_eliminator, elab_strategy] +@[elab_as_eliminator] def last_cases {n : ℕ} {C : fin (n + 1) → Sort*} (hlast : C (fin.last n)) (hcast : (Π (i : fin n), C i.cast_succ)) (i : fin (n + 1)) : C i := reverse_induction hlast (λ i _, hcast i) i @@ -1254,7 +1254,7 @@ reverse_induction_cast_succ _ _ _ /-- Define `f : Π i : fin (m + n), C i` by separately handling the cases `i = cast_add n i`, `j : fin m` and `i = nat_add m j`, `j : fin n`. -/ -@[elab_as_eliminator, elab_strategy] +@[elab_as_eliminator] def add_cases {m n : ℕ} {C : fin (m + n) → Sort u} (hleft : Π i, C (cast_add n i)) (hright : Π i, C (nat_add m i)) (i : fin (m + n)) : C i := @@ -1263,7 +1263,7 @@ else eq.rec_on (nat_add_sub_nat_cast (le_of_not_lt hi)) (hright _) @[simp] lemma add_cases_left {m n : ℕ} {C : fin (m + n) → Sort*} (hleft : Π i, C (cast_add n i)) (hright : Π i, C (nat_add m i)) (i : fin m) : - add_cases hleft hright (fin.cast_add n i) = hleft i := + @add_cases _ _ C hleft hright (fin.cast_add n i) = hleft i := begin cases i with i hi, rw [add_cases, dif_pos (cast_add_lt _ _)], @@ -1272,7 +1272,7 @@ end @[simp] lemma add_cases_right {m n : ℕ} {C : fin (m + n) → Sort*} (hleft : Π i, C (cast_add n i)) (hright : Π i, C (nat_add m i)) (i : fin n) : - add_cases hleft hright (nat_add m i) = hright i := + @add_cases _ _ C hleft hright (nat_add m i) = hright i := begin have : ¬ (nat_add m i : ℕ) < m, from (le_coe_nat_add _ _).not_lt, rw [add_cases, dif_neg this], From d3e8e0a0237c10c2627bf52c246b15ff8e7df4c0 Mon Sep 17 00:00:00 2001 From: Paul Lezeau <72892199+Paul-Lez@users.noreply.github.com> Date: Wed, 4 Jan 2023 00:19:47 +0000 Subject: [PATCH 0139/1291] chore(field_theory/minpoly): Split results from `minpoly.lean` into three files (#18048) In this PR, we split the results contained in `minpoly.lean` into three separate files: `minpoly.basic.lean`, `minpoly.field.lean` and `minpoly.gcd_monoid.lean`. This is helpful for PR #18021. --- src/field_theory/intermediate_field.lean | 2 +- src/field_theory/minpoly.lean | 580 ------------------ src/field_theory/minpoly/basic.lean | 253 ++++++++ src/field_theory/minpoly/default.lean | 3 + src/field_theory/minpoly/field.lean | 256 ++++++++ src/field_theory/minpoly/gcd_monoid.lean | 124 ++++ src/field_theory/separable.lean | 2 +- .../annihilating_polynomial.lean | 2 +- src/linear_algebra/charpoly/basic.lean | 2 +- src/ring_theory/adjoin_root.lean | 1 + src/ring_theory/power_basis.lean | 2 +- 11 files changed, 642 insertions(+), 585 deletions(-) delete mode 100644 src/field_theory/minpoly.lean create mode 100644 src/field_theory/minpoly/basic.lean create mode 100644 src/field_theory/minpoly/default.lean create mode 100644 src/field_theory/minpoly/field.lean create mode 100644 src/field_theory/minpoly/gcd_monoid.lean diff --git a/src/field_theory/intermediate_field.lean b/src/field_theory/intermediate_field.lean index 5908a0709bfed..8d184914aea9c 100644 --- a/src/field_theory/intermediate_field.lean +++ b/src/field_theory/intermediate_field.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Anne Baanen -/ -import field_theory.minpoly +import field_theory.minpoly.field import field_theory.subfield import field_theory.tower diff --git a/src/field_theory/minpoly.lean b/src/field_theory/minpoly.lean deleted file mode 100644 index 7620941b99e77..0000000000000 --- a/src/field_theory/minpoly.lean +++ /dev/null @@ -1,580 +0,0 @@ -/- -Copyright (c) 2019 Chris Hughes. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Chris Hughes, Johan Commelin --/ -import data.polynomial.field_division -import ring_theory.integral_closure -import ring_theory.polynomial.gauss_lemma - -/-! -# Minimal polynomials - -This file defines the minimal polynomial of an element `x` of an `A`-algebra `B`, -under the assumption that x is integral over `A`. - -After stating the defining property we specialize to the setting of field extensions -and derive some well-known properties, amongst which the fact that minimal polynomials -are irreducible, and uniquely determined by their defining property. - --/ - -open_locale classical polynomial -open polynomial set function - -variables {A B : Type*} - -section min_poly_def -variables (A) [comm_ring A] [ring B] [algebra A B] - -/-- -Suppose `x : B`, where `B` is an `A`-algebra. - -The minimal polynomial `minpoly A x` of `x` -is a monic polynomial with coefficients in `A` of smallest degree that has `x` as its root, -if such exists (`is_integral A x`) or zero otherwise. - -For example, if `V` is a `𝕜`-vector space for some field `𝕜` and `f : V →ₗ[𝕜] V` then -the minimal polynomial of `f` is `minpoly 𝕜 f`. --/ -noncomputable def minpoly (x : B) : A[X] := -if hx : is_integral A x then well_founded.min degree_lt_wf _ hx else 0 - -end min_poly_def - -namespace minpoly - -section ring -variables [comm_ring A] [ring B] [algebra A B] -variables {x : B} - -/-- A minimal polynomial is monic. -/ -lemma monic (hx : is_integral A x) : monic (minpoly A x) := -by { delta minpoly, rw dif_pos hx, exact (well_founded.min_mem degree_lt_wf _ hx).1 } - -/-- A minimal polynomial is nonzero. -/ -lemma ne_zero [nontrivial A] (hx : is_integral A x) : minpoly A x ≠ 0 := -(monic hx).ne_zero - -lemma eq_zero (hx : ¬ is_integral A x) : minpoly A x = 0 := -dif_neg hx - -variables (A x) - -/-- An element is a root of its minimal polynomial. -/ -@[simp] lemma aeval : aeval x (minpoly A x) = 0 := -begin - delta minpoly, split_ifs with hx, - { exact (well_founded.min_mem degree_lt_wf _ hx).2 }, - { exact aeval_zero _ } -end - -/-- A minimal polynomial is not `1`. -/ -lemma ne_one [nontrivial B] : minpoly A x ≠ 1 := -begin - intro h, - refine (one_ne_zero : (1 : B) ≠ 0) _, - simpa using congr_arg (polynomial.aeval x) h -end - -lemma map_ne_one [nontrivial B] {R : Type*} [semiring R] [nontrivial R] (f : A →+* R) : - (minpoly A x).map f ≠ 1 := -begin - by_cases hx : is_integral A x, - { exact mt ((monic hx).eq_one_of_map_eq_one f) (ne_one A x) }, - { rw [eq_zero hx, polynomial.map_zero], exact zero_ne_one }, -end - -/-- A minimal polynomial is not a unit. -/ -lemma not_is_unit [nontrivial B] : ¬ is_unit (minpoly A x) := -begin - haveI : nontrivial A := (algebra_map A B).domain_nontrivial, - by_cases hx : is_integral A x, - { exact mt (monic hx).eq_one_of_is_unit (ne_one A x) }, - { rw [eq_zero hx], exact not_is_unit_zero } -end - -lemma mem_range_of_degree_eq_one (hx : (minpoly A x).degree = 1) : x ∈ (algebra_map A B).range := -begin - have h : is_integral A x, - { by_contra h, - rw [eq_zero h, degree_zero, ←with_bot.coe_one] at hx, - exact (ne_of_lt (show ⊥ < ↑1, from with_bot.bot_lt_coe 1) hx) }, - have key := minpoly.aeval A x, - rw [eq_X_add_C_of_degree_eq_one hx, (minpoly.monic h).leading_coeff, C_1, one_mul, aeval_add, - aeval_C, aeval_X, ←eq_neg_iff_add_eq_zero, ←ring_hom.map_neg] at key, - exact ⟨-(minpoly A x).coeff 0, key.symm⟩, -end - -/-- The defining property of the minimal polynomial of an element `x`: -it is the monic polynomial with smallest degree that has `x` as its root. -/ -lemma min {p : A[X]} (pmonic : p.monic) (hp : polynomial.aeval x p = 0) : - degree (minpoly A x) ≤ degree p := -begin - delta minpoly, split_ifs with hx, - { exact le_of_not_lt (well_founded.not_lt_min degree_lt_wf _ hx ⟨pmonic, hp⟩) }, - { simp only [degree_zero, bot_le] } -end - -@[nontriviality] lemma subsingleton [subsingleton B] : minpoly A x = 1 := -begin - nontriviality A, - have := minpoly.min A x monic_one (subsingleton.elim _ _), - rw degree_one at this, - cases le_or_lt (minpoly A x).degree 0 with h h, - { rwa (monic ⟨1, monic_one, by simp⟩ : (minpoly A x).monic).degree_le_zero_iff_eq_one at h }, - { exact (this.not_lt h).elim }, -end - -end ring - -section comm_ring - -variables [comm_ring A] - -section ring - -variables [ring B] [algebra A B] [nontrivial B] -variables {x : B} - -/-- The degree of a minimal polynomial, as a natural number, is positive. -/ -lemma nat_degree_pos (hx : is_integral A x) : 0 < nat_degree (minpoly A x) := -begin - rw pos_iff_ne_zero, - intro ndeg_eq_zero, - have eq_one : minpoly A x = 1, - { rw eq_C_of_nat_degree_eq_zero ndeg_eq_zero, convert C_1, - simpa only [ndeg_eq_zero.symm] using (monic hx).leading_coeff }, - simpa only [eq_one, alg_hom.map_one, one_ne_zero] using aeval A x -end - -/-- The degree of a minimal polynomial is positive. -/ -lemma degree_pos (hx : is_integral A x) : 0 < degree (minpoly A x) := -nat_degree_pos_iff_degree_pos.mp (nat_degree_pos hx) - -/-- If `B/A` is an injective ring extension, and `a` is an element of `A`, -then the minimal polynomial of `algebra_map A B a` is `X - C a`. -/ -lemma eq_X_sub_C_of_algebra_map_inj - (a : A) (hf : function.injective (algebra_map A B)) : - minpoly A (algebra_map A B a) = X - C a := -begin - nontriviality A, - have hdegle : (minpoly A (algebra_map A B a)).nat_degree ≤ 1, - { apply with_bot.coe_le_coe.1, - rw [←degree_eq_nat_degree (ne_zero (@is_integral_algebra_map A B _ _ _ a)), - with_top.coe_one, ←degree_X_sub_C a], - refine min A (algebra_map A B a) (monic_X_sub_C a) _, - simp only [aeval_C, aeval_X, alg_hom.map_sub, sub_self] }, - have hdeg : (minpoly A (algebra_map A B a)).degree = 1, - { apply (degree_eq_iff_nat_degree_eq (ne_zero (@is_integral_algebra_map A B _ _ _ a))).2, - apply le_antisymm hdegle (nat_degree_pos (@is_integral_algebra_map A B _ _ _ a)) }, - have hrw := eq_X_add_C_of_degree_eq_one hdeg, - simp only [monic (@is_integral_algebra_map A B _ _ _ a), one_mul, - monic.leading_coeff, ring_hom.map_one] at hrw, - have h0 : (minpoly A (algebra_map A B a)).coeff 0 = -a, - { have hroot := aeval A (algebra_map A B a), - rw [hrw, add_comm] at hroot, - simp only [aeval_C, aeval_X, aeval_add] at hroot, - replace hroot := eq_neg_of_add_eq_zero_left hroot, - rw [←ring_hom.map_neg _ a] at hroot, - exact (hf hroot) }, - rw hrw, - simp only [h0, ring_hom.map_neg, sub_eq_add_neg], -end - -end ring - -section is_domain - -variables [is_domain A] [ring B] [algebra A B] -variables {x : B} - -/-- If `a` strictly divides the minimal polynomial of `x`, then `x` cannot be a root for `a`. -/ -lemma aeval_ne_zero_of_dvd_not_unit_minpoly {a : A[X]} (hx : is_integral A x) - (hamonic : a.monic) (hdvd : dvd_not_unit a (minpoly A x)) : - polynomial.aeval x a ≠ 0 := -begin - intro ha, - refine not_lt_of_ge (minpoly.min A x hamonic ha) _, - obtain ⟨hzeroa, b, hb_nunit, prod⟩ := hdvd, - have hbmonic : b.monic, - { rw monic.def, - have := monic hx, - rwa [monic.def, prod, leading_coeff_mul, monic.def.mp hamonic, one_mul] at this }, - have hzerob : b ≠ 0 := hbmonic.ne_zero, - have degbzero : 0 < b.nat_degree, - { apply nat.pos_of_ne_zero, - intro h, - have h₁ := eq_C_of_nat_degree_eq_zero h, - rw [←h, ←leading_coeff, monic.def.1 hbmonic, C_1] at h₁, - rw h₁ at hb_nunit, - have := is_unit_one, - contradiction }, - rw [prod, degree_mul, degree_eq_nat_degree hzeroa, degree_eq_nat_degree hzerob], - exact_mod_cast lt_add_of_pos_right _ degbzero, -end - -variables [is_domain B] - -/-- A minimal polynomial is irreducible. -/ -lemma irreducible (hx : is_integral A x) : irreducible (minpoly A x) := -begin - cases irreducible_or_factor (minpoly A x) (not_is_unit A x) with hirr hred, - { exact hirr }, - exfalso, - obtain ⟨a, b, ha_nunit, hb_nunit, hab_eq⟩ := hred, - have coeff_prod : a.leading_coeff * b.leading_coeff = 1, - { rw [←monic.def.1 (monic hx), ←hab_eq], - simp only [leading_coeff_mul] }, - have hamonic : (a * C b.leading_coeff).monic, - { rw monic.def, - simp only [coeff_prod, leading_coeff_mul, leading_coeff_C] }, - have hbmonic : (b * C a.leading_coeff).monic, - { rw [monic.def, mul_comm], - simp only [coeff_prod, leading_coeff_mul, leading_coeff_C] }, - have prod : minpoly A x = (a * C b.leading_coeff) * (b * C a.leading_coeff), - { symmetry, - calc a * C b.leading_coeff * (b * C a.leading_coeff) - = a * b * (C a.leading_coeff * C b.leading_coeff) : by ring - ... = a * b * (C (a.leading_coeff * b.leading_coeff)) : by simp only [ring_hom.map_mul] - ... = a * b : by rw [coeff_prod, C_1, mul_one] - ... = minpoly A x : hab_eq }, - have hzero := aeval A x, - rw [prod, aeval_mul, mul_eq_zero] at hzero, - cases hzero, - { refine aeval_ne_zero_of_dvd_not_unit_minpoly hx hamonic _ hzero, - exact ⟨hamonic.ne_zero, _, mt is_unit_of_mul_is_unit_left hb_nunit, prod⟩ }, - { refine aeval_ne_zero_of_dvd_not_unit_minpoly hx hbmonic _ hzero, - rw mul_comm at prod, - exact ⟨hbmonic.ne_zero, _, mt is_unit_of_mul_is_unit_left ha_nunit, prod⟩ }, -end - -end is_domain - -end comm_ring - -section field -variables [field A] - -section ring -variables [ring B] [algebra A B] -variables {x : B} - -variables (A x) - -/-- If an element `x` is a root of a nonzero polynomial `p`, then the degree of `p` is at least the -degree of the minimal polynomial of `x`. See also `gcd_domain_degree_le_of_ne_zero` which relaxes -the assumptions on `A` in exchange for stronger assumptions on `B`. -/ -lemma degree_le_of_ne_zero - {p : A[X]} (pnz : p ≠ 0) (hp : polynomial.aeval x p = 0) : - degree (minpoly A x) ≤ degree p := -calc degree (minpoly A x) ≤ degree (p * C (leading_coeff p)⁻¹) : - min A x (monic_mul_leading_coeff_inv pnz) (by simp [hp]) - ... = degree p : degree_mul_leading_coeff_inv p pnz - -lemma ne_zero_of_finite_field_extension (e : B) [finite_dimensional A B] : minpoly A e ≠ 0 := -minpoly.ne_zero $ is_integral_of_noetherian (is_noetherian.iff_fg.2 infer_instance) _ - -/-- The minimal polynomial of an element `x` is uniquely characterized by its defining property: -if there is another monic polynomial of minimal degree that has `x` as a root, then this polynomial -is equal to the minimal polynomial of `x`. See also `minpoly.gcd_unique` which relaxes the -assumptions on `A` in exchange for stronger assumptions on `B`. -/ -lemma unique {p : A[X]} - (pmonic : p.monic) (hp : polynomial.aeval x p = 0) - (pmin : ∀ q : A[X], q.monic → polynomial.aeval x q = 0 → degree p ≤ degree q) : - p = minpoly A x := -begin - have hx : is_integral A x := ⟨p, pmonic, hp⟩, - symmetry, apply eq_of_sub_eq_zero, - by_contra hnz, - have := degree_le_of_ne_zero A x hnz (by simp [hp]), - contrapose! this, - apply degree_sub_lt _ (ne_zero hx), - { rw [(monic hx).leading_coeff, pmonic.leading_coeff] }, - { exact le_antisymm (min A x pmonic hp) - (pmin (minpoly A x) (monic hx) (aeval A x)) } -end - -/-- If an element `x` is a root of a polynomial `p`, then the minimal polynomial of `x` divides `p`. -See also `minpoly.gcd_domain_dvd` which relaxes the assumptions on `A` in exchange for stronger -assumptions on `B`. -/ -lemma dvd {p : A[X]} (hp : polynomial.aeval x p = 0) : minpoly A x ∣ p := -begin - by_cases hp0 : p = 0, - { simp only [hp0, dvd_zero] }, - have hx : is_integral A x, - { rw ← is_algebraic_iff_is_integral, exact ⟨p, hp0, hp⟩ }, - rw ← dvd_iff_mod_by_monic_eq_zero (monic hx), - by_contra hnz, - have := degree_le_of_ne_zero A x hnz _, - { contrapose! this, - exact degree_mod_by_monic_lt _ (monic hx) }, - { rw ← mod_by_monic_add_div p (monic hx) at hp, - simpa using hp } -end - -lemma dvd_map_of_is_scalar_tower (A K : Type*) {R : Type*} [comm_ring A] [field K] [comm_ring R] - [algebra A K] [algebra A R] [algebra K R] [is_scalar_tower A K R] (x : R) : - minpoly K x ∣ (minpoly A x).map (algebra_map A K) := -by { refine minpoly.dvd K x _, rw [aeval_map_algebra_map, minpoly.aeval] } - -/-- If `y` is a conjugate of `x` over a field `K`, then it is a conjugate over a subring `R`. -/ -lemma aeval_of_is_scalar_tower (R : Type*) {K T U : Type*} [comm_ring R] [field K] [comm_ring T] - [algebra R K] [algebra K T] [algebra R T] [is_scalar_tower R K T] - [comm_semiring U] [algebra K U] [algebra R U] [is_scalar_tower R K U] - (x : T) (y : U) - (hy : polynomial.aeval y (minpoly K x) = 0) : polynomial.aeval y (minpoly R x) = 0 := -aeval_map_algebra_map K y (minpoly R x) ▸ eval₂_eq_zero_of_dvd_of_eval₂_eq_zero (algebra_map K U) - y (minpoly.dvd_map_of_is_scalar_tower R K x) hy - -variables {A x} - -theorem eq_of_irreducible_of_monic - [nontrivial B] {p : A[X]} (hp1 : _root_.irreducible p) - (hp2 : polynomial.aeval x p = 0) (hp3 : p.monic) : p = minpoly A x := -let ⟨q, hq⟩ := dvd A x hp2 in -eq_of_monic_of_associated hp3 (monic ⟨p, ⟨hp3, hp2⟩⟩) $ -mul_one (minpoly A x) ▸ hq.symm ▸ associated.mul_left _ $ -associated_one_iff_is_unit.2 $ (hp1.is_unit_or_is_unit hq).resolve_left $ not_is_unit A x - -lemma eq_of_irreducible [nontrivial B] {p : A[X]} - (hp1 : _root_.irreducible p) (hp2 : polynomial.aeval x p = 0) : - p * C p.leading_coeff⁻¹ = minpoly A x := -begin - have : p.leading_coeff ≠ 0 := leading_coeff_ne_zero.mpr hp1.ne_zero, - apply eq_of_irreducible_of_monic, - { exact associated.irreducible ⟨⟨C p.leading_coeff⁻¹, C p.leading_coeff, - by rwa [←C_mul, inv_mul_cancel, C_1], by rwa [←C_mul, mul_inv_cancel, C_1]⟩, rfl⟩ hp1 }, - { rw [aeval_mul, hp2, zero_mul] }, - { rwa [polynomial.monic, leading_coeff_mul, leading_coeff_C, mul_inv_cancel] }, -end - -/-- If `y` is the image of `x` in an extension, their minimal polynomials coincide. - -We take `h : y = algebra_map L T x` as an argument because `rw h` typically fails -since `is_integral R y` depends on y. --/ -lemma eq_of_algebra_map_eq {K S T : Type*} [field K] [comm_ring S] [comm_ring T] - [algebra K S] [algebra K T] [algebra S T] - [is_scalar_tower K S T] (hST : function.injective (algebra_map S T)) - {x : S} {y : T} (hx : is_integral K x) (h : y = algebra_map S T x) : - minpoly K x = minpoly K y := -minpoly.unique _ _ (minpoly.monic hx) - (by rw [h, aeval_algebra_map_apply, minpoly.aeval, ring_hom.map_zero]) - (λ q q_monic root_q, minpoly.min _ _ q_monic - ((aeval_algebra_map_eq_zero_iff_of_injective hST).mp - (h ▸ root_q : polynomial.aeval (algebra_map S T x) q = 0))) - -lemma add_algebra_map {B : Type*} [comm_ring B] [algebra A B] {x : B} - (hx : is_integral A x) (a : A) : - minpoly A (x + (algebra_map A B a)) = (minpoly A x).comp (X - C a) := -begin - refine (minpoly.unique _ _ ((minpoly.monic hx).comp_X_sub_C _) _ (λ q qmo hq, _)).symm, - { simp [aeval_comp] }, - { have : (polynomial.aeval x) (q.comp (X + C a)) = 0 := by simpa [aeval_comp] using hq, - have H := minpoly.min A x (qmo.comp_X_add_C _) this, - rw [degree_eq_nat_degree qmo.ne_zero, degree_eq_nat_degree - ((minpoly.monic hx).comp_X_sub_C _).ne_zero, with_bot.coe_le_coe, nat_degree_comp, - nat_degree_X_sub_C, mul_one], - rwa [degree_eq_nat_degree (minpoly.ne_zero hx), degree_eq_nat_degree - (qmo.comp_X_add_C _).ne_zero, with_bot.coe_le_coe, nat_degree_comp, - nat_degree_X_add_C, mul_one] at H } -end - -lemma sub_algebra_map {B : Type*} [comm_ring B] [algebra A B] {x : B} - (hx : is_integral A x) (a : A) : - minpoly A (x - (algebra_map A B a)) = (minpoly A x).comp (X + C a) := -by simpa [sub_eq_add_neg] using add_algebra_map hx (-a) - -section alg_hom_fintype - -/-- A technical finiteness result. -/ -noncomputable def fintype.subtype_prod {E : Type*} {X : set E} (hX : X.finite) {L : Type*} - (F : E → multiset L) : fintype (Π x : X, {l : L // l ∈ F x}) := -let hX := finite.fintype hX in by exactI pi.fintype - -variables (F E K : Type*) [field F] [ring E] [comm_ring K] [is_domain K] - [algebra F E] [algebra F K] [finite_dimensional F E] - -/-- Function from Hom_K(E,L) to pi type Π (x : basis), roots of min poly of x -/ --- Marked as `noncomputable!` since this definition takes multiple seconds to compile, --- and isn't very computable in practice (since neither `finrank` nor `fin_basis` are). -noncomputable! def roots_of_min_poly_pi_type (φ : E →ₐ[F] K) - (x : range (finite_dimensional.fin_basis F E : _ → E)) : - {l : K // l ∈ (((minpoly F x.1).map (algebra_map F K)).roots : multiset K)} := -⟨φ x, by rw [mem_roots_map (minpoly.ne_zero_of_finite_field_extension F x.val), - subtype.val_eq_coe, ←aeval_def, aeval_alg_hom_apply, minpoly.aeval, map_zero]⟩ - -lemma aux_inj_roots_of_min_poly : injective (roots_of_min_poly_pi_type F E K) := -begin - intros f g h, - suffices : (f : E →ₗ[F] K) = g, - { rwa fun_like.ext'_iff at this ⊢ }, - rw funext_iff at h, - exact linear_map.ext_on (finite_dimensional.fin_basis F E).span_eq - (λ e he, subtype.ext_iff.mp (h ⟨e, he⟩)), -end - -/-- Given field extensions `E/F` and `K/F`, with `E/F` finite, there are finitely many `F`-algebra - homomorphisms `E →ₐ[K] K`. -/ -noncomputable instance alg_hom.fintype : fintype (E →ₐ[F] K) := -@fintype.of_injective _ _ (fintype.subtype_prod (finite_range (finite_dimensional.fin_basis F E)) - (λ e, ((minpoly F e).map (algebra_map F K)).roots)) _ (aux_inj_roots_of_min_poly F E K) - -end alg_hom_fintype - -section gcd_domain - -variables {R S : Type*} (K L : Type*) [comm_ring R] [is_domain R] [normalized_gcd_monoid R] - [field K] [comm_ring S] [is_domain S] [algebra R K] [is_fraction_ring R K] [algebra R S] [field L] - [algebra S L] [algebra K L] [algebra R L] [is_scalar_tower R K L] [is_scalar_tower R S L] - {s : S} (hs : is_integral R s) - -include hs - -/-- For GCD domains, the minimal polynomial over the ring is the same as the minimal polynomial -over the fraction field. See `minpoly.gcd_domain_eq_field_fractions'` if `S` is already a -`K`-algebra. -/ -lemma gcd_domain_eq_field_fractions : - minpoly K (algebra_map S L s) = (minpoly R s).map (algebra_map R K) := -begin - refine (eq_of_irreducible_of_monic _ _ _).symm, - { exact (polynomial.is_primitive.irreducible_iff_irreducible_map_fraction_map - (polynomial.monic.is_primitive (monic hs))).1 (irreducible hs) }, - { rw [aeval_map_algebra_map, aeval_algebra_map_apply, aeval, map_zero] }, - { exact (monic hs).map _ } -end - -/-- For GCD domains, the minimal polynomial over the ring is the same as the minimal polynomial -over the fraction field. Compared to `minpoly.gcd_domain_eq_field_fractions`, this version is useful -if the element is in a ring that is already a `K`-algebra. -/ -lemma gcd_domain_eq_field_fractions' [algebra K S] [is_scalar_tower R K S] : - minpoly K s = (minpoly R s).map (algebra_map R K) := -begin - let L := fraction_ring S, - rw [← gcd_domain_eq_field_fractions K L hs], - refine minpoly.eq_of_algebra_map_eq (is_fraction_ring.injective S L) - (is_integral_of_is_scalar_tower hs) rfl -end - -variable [no_zero_smul_divisors R S] - -/-- For GCD domains, the minimal polynomial divides any primitive polynomial that has the integral -element as root. See also `minpoly.dvd` which relaxes the assumptions on `S` in exchange for -stronger assumptions on `R`. -/ -lemma gcd_domain_dvd {P : R[X]} (hP : P ≠ 0) (hroot : polynomial.aeval s P = 0) : minpoly R s ∣ P := -begin - let K := fraction_ring R, - let L := fraction_ring S, - let P₁ := P.prim_part, - suffices : minpoly R s ∣ P₁, - { exact dvd_trans this (prim_part_dvd _) }, - apply (is_primitive.dvd_iff_fraction_map_dvd_fraction_map K (monic hs).is_primitive - P.is_primitive_prim_part).2, - let y := algebra_map S L s, - have hy : is_integral R y := hs.algebra_map, - rw [← gcd_domain_eq_field_fractions K L hs], - refine dvd _ _ _, - rw [aeval_map_algebra_map, aeval_algebra_map_apply, aeval_prim_part_eq_zero hP hroot, map_zero] -end - -/-- If an element `x` is a root of a nonzero polynomial `p`, then the degree of `p` is at least the -degree of the minimal polynomial of `x`. See also `minpoly.degree_le_of_ne_zero` which relaxes the -assumptions on `S` in exchange for stronger assumptions on `R`. -/ -lemma gcd_domain_degree_le_of_ne_zero {p : R[X]} (hp0 : p ≠ 0) (hp : polynomial.aeval s p = 0) : - degree (minpoly R s) ≤ degree p := -begin - rw [degree_eq_nat_degree (minpoly.ne_zero hs), degree_eq_nat_degree hp0], - norm_cast, - exact nat_degree_le_of_dvd (gcd_domain_dvd hs hp0 hp) hp0 -end - -omit hs - -/-- The minimal polynomial of an element `x` is uniquely characterized by its defining property: -if there is another monic polynomial of minimal degree that has `x` as a root, then this polynomial -is equal to the minimal polynomial of `x`. See also `minpoly.unique` which relaxes the -assumptions on `S` in exchange for stronger assumptions on `R`. -/ -lemma gcd_domain_unique {P : R[X]} (hmo : P.monic) (hP : polynomial.aeval s P = 0) - (Pmin : ∀ Q : R[X], Q.monic → polynomial.aeval s Q = 0 → degree P ≤ degree Q) : - P = minpoly R s := -begin - have hs : is_integral R s := ⟨P, hmo, hP⟩, - symmetry, apply eq_of_sub_eq_zero, - by_contra hnz, - have := gcd_domain_degree_le_of_ne_zero hs hnz (by simp [hP]), - contrapose! this, - refine degree_sub_lt _ (ne_zero hs) _, - { exact le_antisymm (min R s hmo hP) - (Pmin (minpoly R s) (monic hs) (aeval R s)) }, - { rw [(monic hs).leading_coeff, hmo.leading_coeff] } -end - -end gcd_domain - -variables (B) [nontrivial B] - -/-- If `B/K` is a nontrivial algebra over a field, and `x` is an element of `K`, -then the minimal polynomial of `algebra_map K B x` is `X - C x`. -/ -lemma eq_X_sub_C (a : A) : minpoly A (algebra_map A B a) = X - C a := -eq_X_sub_C_of_algebra_map_inj a (algebra_map A B).injective - -lemma eq_X_sub_C' (a : A) : minpoly A a = X - C a := eq_X_sub_C A a - -variables (A) - -/-- The minimal polynomial of `0` is `X`. -/ -@[simp] lemma zero : minpoly A (0:B) = X := -by simpa only [add_zero, C_0, sub_eq_add_neg, neg_zero, ring_hom.map_zero] - using eq_X_sub_C B (0:A) - -/-- The minimal polynomial of `1` is `X - 1`. -/ -@[simp] lemma one : minpoly A (1:B) = X - 1 := -by simpa only [ring_hom.map_one, C_1, sub_eq_add_neg] using eq_X_sub_C B (1:A) - -end ring - -section is_domain -variables [ring B] [is_domain B] [algebra A B] -variables {x : B} - -/-- A minimal polynomial is prime. -/ -lemma prime (hx : is_integral A x) : prime (minpoly A x) := -begin - refine ⟨ne_zero hx, not_is_unit A x, _⟩, - rintros p q ⟨d, h⟩, - have : polynomial.aeval x (p*q) = 0 := by simp [h, aeval A x], - replace : polynomial.aeval x p = 0 ∨ polynomial.aeval x q = 0 := by simpa, - exact or.imp (dvd A x) (dvd A x) this -end - -/-- If `L/K` is a field extension and an element `y` of `K` is a root of the minimal polynomial -of an element `x ∈ L`, then `y` maps to `x` under the field embedding. -/ -lemma root {x : B} (hx : is_integral A x) {y : A} (h : is_root (minpoly A x) y) : - algebra_map A B y = x := -have key : minpoly A x = X - C y := -eq_of_monic_of_associated (monic hx) (monic_X_sub_C y) (associated_of_dvd_dvd - ((irreducible_X_sub_C y).dvd_symm (irreducible hx) (dvd_iff_is_root.2 h)) - (dvd_iff_is_root.2 h)), -by { have := aeval A x, rwa [key, alg_hom.map_sub, aeval_X, aeval_C, sub_eq_zero, eq_comm] at this } - -/-- The constant coefficient of the minimal polynomial of `x` is `0` if and only if `x = 0`. -/ -@[simp] lemma coeff_zero_eq_zero (hx : is_integral A x) : coeff (minpoly A x) 0 = 0 ↔ x = 0 := -begin - split, - { intro h, - have zero_root := zero_is_root_of_coeff_zero_eq_zero h, - rw ← root hx zero_root, - exact ring_hom.map_zero _ }, - { rintro rfl, simp } -end - -/-- The minimal polynomial of a nonzero element has nonzero constant coefficient. -/ -lemma coeff_zero_ne_zero (hx : is_integral A x) (h : x ≠ 0) : coeff (minpoly A x) 0 ≠ 0 := -by { contrapose! h, simpa only [hx, coeff_zero_eq_zero] using h } - -end is_domain - -end field - -end minpoly diff --git a/src/field_theory/minpoly/basic.lean b/src/field_theory/minpoly/basic.lean new file mode 100644 index 0000000000000..529dc70c0c938 --- /dev/null +++ b/src/field_theory/minpoly/basic.lean @@ -0,0 +1,253 @@ +/- +Copyright (c) 2019 Chris Hughes. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Chris Hughes, Johan Commelin +-/ +import data.polynomial.field_division +import ring_theory.integral_closure +import ring_theory.polynomial.gauss_lemma + +/-! +# Minimal polynomials + +This file defines the minimal polynomial of an element `x` of an `A`-algebra `B`, +under the assumption that x is integral over `A`, and derives some basic properties +such as ireducibility under the assumption `B` is a domain. + +-/ + +open_locale classical polynomial +open polynomial set function + +variables {A B : Type*} + +section min_poly_def +variables (A) [comm_ring A] [ring B] [algebra A B] + +/-- +Suppose `x : B`, where `B` is an `A`-algebra. + +The minimal polynomial `minpoly A x` of `x` +is a monic polynomial with coefficients in `A` of smallest degree that has `x` as its root, +if such exists (`is_integral A x`) or zero otherwise. + +For example, if `V` is a `𝕜`-vector space for some field `𝕜` and `f : V →ₗ[𝕜] V` then +the minimal polynomial of `f` is `minpoly 𝕜 f`. +-/ +noncomputable def minpoly (x : B) : A[X] := +if hx : is_integral A x then well_founded.min degree_lt_wf _ hx else 0 + +end min_poly_def + +namespace minpoly + +section ring +variables [comm_ring A] [ring B] [algebra A B] +variables {x : B} + +/-- A minimal polynomial is monic. -/ +lemma monic (hx : is_integral A x) : monic (minpoly A x) := +by { delta minpoly, rw dif_pos hx, exact (well_founded.min_mem degree_lt_wf _ hx).1 } + +/-- A minimal polynomial is nonzero. -/ +lemma ne_zero [nontrivial A] (hx : is_integral A x) : minpoly A x ≠ 0 := +(monic hx).ne_zero + +lemma eq_zero (hx : ¬ is_integral A x) : minpoly A x = 0 := +dif_neg hx + +variables (A x) + +/-- An element is a root of its minimal polynomial. -/ +@[simp] lemma aeval : aeval x (minpoly A x) = 0 := +begin + delta minpoly, split_ifs with hx, + { exact (well_founded.min_mem degree_lt_wf _ hx).2 }, + { exact aeval_zero _ } +end + +/-- A minimal polynomial is not `1`. -/ +lemma ne_one [nontrivial B] : minpoly A x ≠ 1 := +begin + intro h, + refine (one_ne_zero : (1 : B) ≠ 0) _, + simpa using congr_arg (polynomial.aeval x) h +end + +lemma map_ne_one [nontrivial B] {R : Type*} [semiring R] [nontrivial R] (f : A →+* R) : + (minpoly A x).map f ≠ 1 := +begin + by_cases hx : is_integral A x, + { exact mt ((monic hx).eq_one_of_map_eq_one f) (ne_one A x) }, + { rw [eq_zero hx, polynomial.map_zero], exact zero_ne_one }, +end + +/-- A minimal polynomial is not a unit. -/ +lemma not_is_unit [nontrivial B] : ¬ is_unit (minpoly A x) := +begin + haveI : nontrivial A := (algebra_map A B).domain_nontrivial, + by_cases hx : is_integral A x, + { exact mt (monic hx).eq_one_of_is_unit (ne_one A x) }, + { rw [eq_zero hx], exact not_is_unit_zero } +end + +lemma mem_range_of_degree_eq_one (hx : (minpoly A x).degree = 1) : x ∈ (algebra_map A B).range := +begin + have h : is_integral A x, + { by_contra h, + rw [eq_zero h, degree_zero, ←with_bot.coe_one] at hx, + exact (ne_of_lt (show ⊥ < ↑1, from with_bot.bot_lt_coe 1) hx) }, + have key := minpoly.aeval A x, + rw [eq_X_add_C_of_degree_eq_one hx, (minpoly.monic h).leading_coeff, C_1, one_mul, aeval_add, + aeval_C, aeval_X, ←eq_neg_iff_add_eq_zero, ←ring_hom.map_neg] at key, + exact ⟨-(minpoly A x).coeff 0, key.symm⟩, +end + +/-- The defining property of the minimal polynomial of an element `x`: +it is the monic polynomial with smallest degree that has `x` as its root. -/ +lemma min {p : A[X]} (pmonic : p.monic) (hp : polynomial.aeval x p = 0) : + degree (minpoly A x) ≤ degree p := +begin + delta minpoly, split_ifs with hx, + { exact le_of_not_lt (well_founded.not_lt_min degree_lt_wf _ hx ⟨pmonic, hp⟩) }, + { simp only [degree_zero, bot_le] } +end + +@[nontriviality] lemma subsingleton [subsingleton B] : minpoly A x = 1 := +begin + nontriviality A, + have := minpoly.min A x monic_one (subsingleton.elim _ _), + rw degree_one at this, + cases le_or_lt (minpoly A x).degree 0 with h h, + { rwa (monic ⟨1, monic_one, by simp⟩ : (minpoly A x).monic).degree_le_zero_iff_eq_one at h }, + { exact (this.not_lt h).elim }, +end + +end ring + +section comm_ring + +variables [comm_ring A] + +section ring + +variables [ring B] [algebra A B] [nontrivial B] +variables {x : B} + +/-- The degree of a minimal polynomial, as a natural number, is positive. -/ +lemma nat_degree_pos (hx : is_integral A x) : 0 < nat_degree (minpoly A x) := +begin + rw pos_iff_ne_zero, + intro ndeg_eq_zero, + have eq_one : minpoly A x = 1, + { rw eq_C_of_nat_degree_eq_zero ndeg_eq_zero, convert C_1, + simpa only [ndeg_eq_zero.symm] using (monic hx).leading_coeff }, + simpa only [eq_one, alg_hom.map_one, one_ne_zero] using aeval A x +end + +/-- The degree of a minimal polynomial is positive. -/ +lemma degree_pos (hx : is_integral A x) : 0 < degree (minpoly A x) := +nat_degree_pos_iff_degree_pos.mp (nat_degree_pos hx) + +/-- If `B/A` is an injective ring extension, and `a` is an element of `A`, +then the minimal polynomial of `algebra_map A B a` is `X - C a`. -/ +lemma eq_X_sub_C_of_algebra_map_inj + (a : A) (hf : function.injective (algebra_map A B)) : + minpoly A (algebra_map A B a) = X - C a := +begin + nontriviality A, + have hdegle : (minpoly A (algebra_map A B a)).nat_degree ≤ 1, + { apply with_bot.coe_le_coe.1, + rw [←degree_eq_nat_degree (ne_zero (@is_integral_algebra_map A B _ _ _ a)), + with_top.coe_one, ←degree_X_sub_C a], + refine min A (algebra_map A B a) (monic_X_sub_C a) _, + simp only [aeval_C, aeval_X, alg_hom.map_sub, sub_self] }, + have hdeg : (minpoly A (algebra_map A B a)).degree = 1, + { apply (degree_eq_iff_nat_degree_eq (ne_zero (@is_integral_algebra_map A B _ _ _ a))).2, + apply le_antisymm hdegle (nat_degree_pos (@is_integral_algebra_map A B _ _ _ a)) }, + have hrw := eq_X_add_C_of_degree_eq_one hdeg, + simp only [monic (@is_integral_algebra_map A B _ _ _ a), one_mul, + monic.leading_coeff, ring_hom.map_one] at hrw, + have h0 : (minpoly A (algebra_map A B a)).coeff 0 = -a, + { have hroot := aeval A (algebra_map A B a), + rw [hrw, add_comm] at hroot, + simp only [aeval_C, aeval_X, aeval_add] at hroot, + replace hroot := eq_neg_of_add_eq_zero_left hroot, + rw [←ring_hom.map_neg _ a] at hroot, + exact (hf hroot) }, + rw hrw, + simp only [h0, ring_hom.map_neg, sub_eq_add_neg], +end + +end ring + +section is_domain + +variables [is_domain A] [ring B] [algebra A B] +variables {x : B} + +/-- If `a` strictly divides the minimal polynomial of `x`, then `x` cannot be a root for `a`. -/ +lemma aeval_ne_zero_of_dvd_not_unit_minpoly {a : A[X]} (hx : is_integral A x) + (hamonic : a.monic) (hdvd : dvd_not_unit a (minpoly A x)) : + polynomial.aeval x a ≠ 0 := +begin + intro ha, + refine not_lt_of_ge (minpoly.min A x hamonic ha) _, + obtain ⟨hzeroa, b, hb_nunit, prod⟩ := hdvd, + have hbmonic : b.monic, + { rw monic.def, + have := monic hx, + rwa [monic.def, prod, leading_coeff_mul, monic.def.mp hamonic, one_mul] at this }, + have hzerob : b ≠ 0 := hbmonic.ne_zero, + have degbzero : 0 < b.nat_degree, + { apply nat.pos_of_ne_zero, + intro h, + have h₁ := eq_C_of_nat_degree_eq_zero h, + rw [←h, ←leading_coeff, monic.def.1 hbmonic, C_1] at h₁, + rw h₁ at hb_nunit, + have := is_unit_one, + contradiction }, + rw [prod, degree_mul, degree_eq_nat_degree hzeroa, degree_eq_nat_degree hzerob], + exact_mod_cast lt_add_of_pos_right _ degbzero, +end + +variables [is_domain B] + +/-- A minimal polynomial is irreducible. -/ +lemma irreducible (hx : is_integral A x) : irreducible (minpoly A x) := +begin + cases irreducible_or_factor (minpoly A x) (not_is_unit A x) with hirr hred, + { exact hirr }, + exfalso, + obtain ⟨a, b, ha_nunit, hb_nunit, hab_eq⟩ := hred, + have coeff_prod : a.leading_coeff * b.leading_coeff = 1, + { rw [←monic.def.1 (monic hx), ←hab_eq], + simp only [leading_coeff_mul] }, + have hamonic : (a * C b.leading_coeff).monic, + { rw monic.def, + simp only [coeff_prod, leading_coeff_mul, leading_coeff_C] }, + have hbmonic : (b * C a.leading_coeff).monic, + { rw [monic.def, mul_comm], + simp only [coeff_prod, leading_coeff_mul, leading_coeff_C] }, + have prod : minpoly A x = (a * C b.leading_coeff) * (b * C a.leading_coeff), + { symmetry, + calc a * C b.leading_coeff * (b * C a.leading_coeff) + = a * b * (C a.leading_coeff * C b.leading_coeff) : by ring + ... = a * b * (C (a.leading_coeff * b.leading_coeff)) : by simp only [ring_hom.map_mul] + ... = a * b : by rw [coeff_prod, C_1, mul_one] + ... = minpoly A x : hab_eq }, + have hzero := aeval A x, + rw [prod, aeval_mul, mul_eq_zero] at hzero, + cases hzero, + { refine aeval_ne_zero_of_dvd_not_unit_minpoly hx hamonic _ hzero, + exact ⟨hamonic.ne_zero, _, mt is_unit_of_mul_is_unit_left hb_nunit, prod⟩ }, + { refine aeval_ne_zero_of_dvd_not_unit_minpoly hx hbmonic _ hzero, + rw mul_comm at prod, + exact ⟨hbmonic.ne_zero, _, mt is_unit_of_mul_is_unit_left ha_nunit, prod⟩ }, +end + +end is_domain + +end comm_ring + +end minpoly diff --git a/src/field_theory/minpoly/default.lean b/src/field_theory/minpoly/default.lean new file mode 100644 index 0000000000000..aa4d2da646584 --- /dev/null +++ b/src/field_theory/minpoly/default.lean @@ -0,0 +1,3 @@ +import field_theory.minpoly.basic +import field_theory.minpoly.field +import field_theory.minpoly.gcd_monoid diff --git a/src/field_theory/minpoly/field.lean b/src/field_theory/minpoly/field.lean new file mode 100644 index 0000000000000..5db46495a2df4 --- /dev/null +++ b/src/field_theory/minpoly/field.lean @@ -0,0 +1,256 @@ +/- +Copyright (c) 2019 Johan Commelin. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Riccardo Brasca, Johan Commelin +-/ +import data.polynomial.field_division +import ring_theory.integral_closure +import field_theory.minpoly.basic + +/-! +# Minimal polynomials on an algebra over a field + +This file specializes the theory of minpoly to the setting of field extensions +and derives some well-known properties, amongst which the fact that minimal polynomials +are irreducible, and uniquely determined by their defining property. + +-/ + +open_locale classical polynomial +open polynomial set function minpoly + +namespace minpoly + +variables {A B : Type*} +variables (A) [field A] + +section ring + +variables [ring B] [algebra A B] (x : B) + +/-- If an element `x` is a root of a nonzero polynomial `p`, then the degree of `p` is at least the +degree of the minimal polynomial of `x`. See also `gcd_domain_degree_le_of_ne_zero` which relaxes +the assumptions on `A` in exchange for stronger assumptions on `B`. -/ +lemma degree_le_of_ne_zero + {p : A[X]} (pnz : p ≠ 0) (hp : polynomial.aeval x p = 0) : + degree (minpoly A x) ≤ degree p := +calc degree (minpoly A x) ≤ degree (p * C (leading_coeff p)⁻¹) : + min A x (monic_mul_leading_coeff_inv pnz) (by simp [hp]) + ... = degree p : degree_mul_leading_coeff_inv p pnz + +lemma ne_zero_of_finite_field_extension (e : B) [finite_dimensional A B] : minpoly A e ≠ 0 := +minpoly.ne_zero $ is_integral_of_noetherian (is_noetherian.iff_fg.2 infer_instance) _ + +/-- The minimal polynomial of an element `x` is uniquely characterized by its defining property: +if there is another monic polynomial of minimal degree that has `x` as a root, then this polynomial +is equal to the minimal polynomial of `x`. See also `minpoly.gcd_unique` which relaxes the +assumptions on `A` in exchange for stronger assumptions on `B`. -/ +lemma unique {p : A[X]} + (pmonic : p.monic) (hp : polynomial.aeval x p = 0) + (pmin : ∀ q : A[X], q.monic → polynomial.aeval x q = 0 → degree p ≤ degree q) : + p = minpoly A x := +begin + have hx : is_integral A x := ⟨p, pmonic, hp⟩, + symmetry, apply eq_of_sub_eq_zero, + by_contra hnz, + have := degree_le_of_ne_zero A x hnz (by simp [hp]), + contrapose! this, + apply degree_sub_lt _ (ne_zero hx), + { rw [(monic hx).leading_coeff, pmonic.leading_coeff] }, + { exact le_antisymm (min A x pmonic hp) + (pmin (minpoly A x) (monic hx) (aeval A x)) } +end + +/-- If an element `x` is a root of a polynomial `p`, then the minimal polynomial of `x` divides `p`. +See also `minpoly.gcd_domain_dvd` which relaxes the assumptions on `A` in exchange for stronger +assumptions on `B`. -/ +lemma dvd {p : A[X]} (hp : polynomial.aeval x p = 0) : minpoly A x ∣ p := +begin + by_cases hp0 : p = 0, + { simp only [hp0, dvd_zero] }, + have hx : is_integral A x, + { rw ← is_algebraic_iff_is_integral, exact ⟨p, hp0, hp⟩ }, + rw ← dvd_iff_mod_by_monic_eq_zero (monic hx), + by_contra hnz, + have := degree_le_of_ne_zero A x hnz _, + { contrapose! this, + exact degree_mod_by_monic_lt _ (monic hx) }, + { rw ← mod_by_monic_add_div p (monic hx) at hp, + simpa using hp } +end + +lemma dvd_map_of_is_scalar_tower (A K : Type*) {R : Type*} [comm_ring A] [field K] [comm_ring R] + [algebra A K] [algebra A R] [algebra K R] [is_scalar_tower A K R] (x : R) : + minpoly K x ∣ (minpoly A x).map (algebra_map A K) := +by { refine minpoly.dvd K x _, rw [aeval_map_algebra_map, minpoly.aeval] } + +/-- If `y` is a conjugate of `x` over a field `K`, then it is a conjugate over a subring `R`. -/ +lemma aeval_of_is_scalar_tower (R : Type*) {K T U : Type*} [comm_ring R] [field K] [comm_ring T] + [algebra R K] [algebra K T] [algebra R T] [is_scalar_tower R K T] + [comm_semiring U] [algebra K U] [algebra R U] [is_scalar_tower R K U] + (x : T) (y : U) + (hy : polynomial.aeval y (minpoly K x) = 0) : polynomial.aeval y (minpoly R x) = 0 := +aeval_map_algebra_map K y (minpoly R x) ▸ eval₂_eq_zero_of_dvd_of_eval₂_eq_zero (algebra_map K U) + y (minpoly.dvd_map_of_is_scalar_tower R K x) hy + +variables {A x} + +theorem eq_of_irreducible_of_monic + [nontrivial B] {p : A[X]} (hp1 : _root_.irreducible p) + (hp2 : polynomial.aeval x p = 0) (hp3 : p.monic) : p = minpoly A x := +let ⟨q, hq⟩ := dvd A x hp2 in +eq_of_monic_of_associated hp3 (monic ⟨p, ⟨hp3, hp2⟩⟩) $ +mul_one (minpoly A x) ▸ hq.symm ▸ associated.mul_left _ $ +associated_one_iff_is_unit.2 $ (hp1.is_unit_or_is_unit hq).resolve_left $ not_is_unit A x + +lemma eq_of_irreducible [nontrivial B] {p : A[X]} + (hp1 : _root_.irreducible p) (hp2 : polynomial.aeval x p = 0) : + p * C p.leading_coeff⁻¹ = minpoly A x := +begin + have : p.leading_coeff ≠ 0 := leading_coeff_ne_zero.mpr hp1.ne_zero, + apply eq_of_irreducible_of_monic, + { exact associated.irreducible ⟨⟨C p.leading_coeff⁻¹, C p.leading_coeff, + by rwa [←C_mul, inv_mul_cancel, C_1], by rwa [←C_mul, mul_inv_cancel, C_1]⟩, rfl⟩ hp1 }, + { rw [aeval_mul, hp2, zero_mul] }, + { rwa [polynomial.monic, leading_coeff_mul, leading_coeff_C, mul_inv_cancel] }, +end + +/-- If `y` is the image of `x` in an extension, their minimal polynomials coincide. + +We take `h : y = algebra_map L T x` as an argument because `rw h` typically fails +since `is_integral R y` depends on y. +-/ +lemma eq_of_algebra_map_eq {K S T : Type*} [field K] [comm_ring S] [comm_ring T] + [algebra K S] [algebra K T] [algebra S T] + [is_scalar_tower K S T] (hST : function.injective (algebra_map S T)) + {x : S} {y : T} (hx : is_integral K x) (h : y = algebra_map S T x) : + minpoly K x = minpoly K y := +minpoly.unique _ _ (minpoly.monic hx) + (by rw [h, aeval_algebra_map_apply, minpoly.aeval, ring_hom.map_zero]) + (λ q q_monic root_q, minpoly.min _ _ q_monic + ((aeval_algebra_map_eq_zero_iff_of_injective hST).mp + (h ▸ root_q : polynomial.aeval (algebra_map S T x) q = 0))) + +lemma add_algebra_map {B : Type*} [comm_ring B] [algebra A B] {x : B} + (hx : is_integral A x) (a : A) : + minpoly A (x + (algebra_map A B a)) = (minpoly A x).comp (X - C a) := +begin + refine (minpoly.unique _ _ ((minpoly.monic hx).comp_X_sub_C _) _ (λ q qmo hq, _)).symm, + { simp [aeval_comp] }, + { have : (polynomial.aeval x) (q.comp (X + C a)) = 0 := by simpa [aeval_comp] using hq, + have H := minpoly.min A x (qmo.comp_X_add_C _) this, + rw [degree_eq_nat_degree qmo.ne_zero, degree_eq_nat_degree + ((minpoly.monic hx).comp_X_sub_C _).ne_zero, with_bot.coe_le_coe, nat_degree_comp, + nat_degree_X_sub_C, mul_one], + rwa [degree_eq_nat_degree (minpoly.ne_zero hx), degree_eq_nat_degree + (qmo.comp_X_add_C _).ne_zero, with_bot.coe_le_coe, nat_degree_comp, + nat_degree_X_add_C, mul_one] at H } +end + +lemma sub_algebra_map {B : Type*} [comm_ring B] [algebra A B] {x : B} + (hx : is_integral A x) (a : A) : + minpoly A (x - (algebra_map A B a)) = (minpoly A x).comp (X + C a) := +by simpa [sub_eq_add_neg] using add_algebra_map hx (-a) + +section alg_hom_fintype + +/-- A technical finiteness result. -/ +noncomputable def fintype.subtype_prod {E : Type*} {X : set E} (hX : X.finite) {L : Type*} + (F : E → multiset L) : fintype (Π x : X, {l : L // l ∈ F x}) := +let hX := finite.fintype hX in by exactI pi.fintype + +variables (F E K : Type*) [field F] [ring E] [comm_ring K] [is_domain K] + [algebra F E] [algebra F K] [finite_dimensional F E] + +/-- Function from Hom_K(E,L) to pi type Π (x : basis), roots of min poly of x -/ +-- Marked as `noncomputable!` since this definition takes multiple seconds to compile, +-- and isn't very computable in practice (since neither `finrank` nor `fin_basis` are). +noncomputable! def roots_of_min_poly_pi_type (φ : E →ₐ[F] K) + (x : range (finite_dimensional.fin_basis F E : _ → E)) : + {l : K // l ∈ (((minpoly F x.1).map (algebra_map F K)).roots : multiset K)} := +⟨φ x, by rw [mem_roots_map (minpoly.ne_zero_of_finite_field_extension F x.val), + subtype.val_eq_coe, ←aeval_def, aeval_alg_hom_apply, minpoly.aeval, map_zero]⟩ + +lemma aux_inj_roots_of_min_poly : injective (roots_of_min_poly_pi_type F E K) := +begin + intros f g h, + suffices : (f : E →ₗ[F] K) = g, + { rwa fun_like.ext'_iff at this ⊢ }, + rw funext_iff at h, + exact linear_map.ext_on (finite_dimensional.fin_basis F E).span_eq + (λ e he, subtype.ext_iff.mp (h ⟨e, he⟩)), +end + +/-- Given field extensions `E/F` and `K/F`, with `E/F` finite, there are finitely many `F`-algebra + homomorphisms `E →ₐ[K] K`. -/ +noncomputable instance alg_hom.fintype : fintype (E →ₐ[F] K) := +@fintype.of_injective _ _ (fintype.subtype_prod (finite_range (finite_dimensional.fin_basis F E)) + (λ e, ((minpoly F e).map (algebra_map F K)).roots)) _ (aux_inj_roots_of_min_poly F E K) + +end alg_hom_fintype + +variables (B) [nontrivial B] + +/-- If `B/K` is a nontrivial algebra over a field, and `x` is an element of `K`, +then the minimal polynomial of `algebra_map K B x` is `X - C x`. -/ +lemma eq_X_sub_C (a : A) : minpoly A (algebra_map A B a) = X - C a := +eq_X_sub_C_of_algebra_map_inj a (algebra_map A B).injective + +lemma eq_X_sub_C' (a : A) : minpoly A a = X - C a := eq_X_sub_C A a + +variables (A) + +/-- The minimal polynomial of `0` is `X`. -/ +@[simp] lemma zero : minpoly A (0:B) = X := +by simpa only [add_zero, C_0, sub_eq_add_neg, neg_zero, ring_hom.map_zero] + using eq_X_sub_C B (0:A) + +/-- The minimal polynomial of `1` is `X - 1`. -/ +@[simp] lemma one : minpoly A (1:B) = X - 1 := +by simpa only [ring_hom.map_one, C_1, sub_eq_add_neg] using eq_X_sub_C B (1:A) + +end ring + +section is_domain + +variables [ring B] [is_domain B] [algebra A B] +variables {A} {x : B} + +/-- A minimal polynomial is prime. -/ +lemma prime (hx : is_integral A x) : prime (minpoly A x) := +begin + refine ⟨ne_zero hx, not_is_unit A x, _⟩, + rintros p q ⟨d, h⟩, + have : polynomial.aeval x (p*q) = 0 := by simp [h, aeval A x], + replace : polynomial.aeval x p = 0 ∨ polynomial.aeval x q = 0 := by simpa, + exact or.imp (dvd A x) (dvd A x) this +end + +/-- If `L/K` is a field extension and an element `y` of `K` is a root of the minimal polynomial +of an element `x ∈ L`, then `y` maps to `x` under the field embedding. -/ +lemma root {x : B} (hx : is_integral A x) {y : A} (h : is_root (minpoly A x) y) : + algebra_map A B y = x := +have key : minpoly A x = X - C y := +eq_of_monic_of_associated (monic hx) (monic_X_sub_C y) (associated_of_dvd_dvd + ((irreducible_X_sub_C y).dvd_symm (irreducible hx) (dvd_iff_is_root.2 h)) + (dvd_iff_is_root.2 h)), +by { have := aeval A x, rwa [key, alg_hom.map_sub, aeval_X, aeval_C, sub_eq_zero, eq_comm] at this } + +/-- The constant coefficient of the minimal polynomial of `x` is `0` if and only if `x = 0`. -/ +@[simp] lemma coeff_zero_eq_zero (hx : is_integral A x) : coeff (minpoly A x) 0 = 0 ↔ x = 0 := +begin + split, + { intro h, + have zero_root := zero_is_root_of_coeff_zero_eq_zero h, + rw ← root hx zero_root, + exact ring_hom.map_zero _ }, + { rintro rfl, simp } +end + +/-- The minimal polynomial of a nonzero element has nonzero constant coefficient. -/ +lemma coeff_zero_ne_zero (hx : is_integral A x) (h : x ≠ 0) : coeff (minpoly A x) 0 ≠ 0 := +by { contrapose! h, simpa only [hx, coeff_zero_eq_zero] using h } + +end is_domain + +end minpoly diff --git a/src/field_theory/minpoly/gcd_monoid.lean b/src/field_theory/minpoly/gcd_monoid.lean new file mode 100644 index 0000000000000..5e56f553bc8a6 --- /dev/null +++ b/src/field_theory/minpoly/gcd_monoid.lean @@ -0,0 +1,124 @@ +/- +Copyright (c) 2019 Riccardo Brasca. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Riccardo Brasca +-/ +import data.polynomial.field_division +import ring_theory.integral_closure +import ring_theory.polynomial.gauss_lemma +import field_theory.minpoly.field + +/-! +# Minimal polynomials over a GCD monoid + +This file specializes the theory of minpoly to the case of an algebra over a GCD monoid. + +## Main results + + * `gcd_domain_eq_field_fractions`: For GCD domains, the minimal polynomial over the ring is the + same as the minimal polynomial over the fraction field. + + * `gcd_domain_dvd` : For GCD domains, the minimal polynomial divides any primitive polynomial + that has the integral element as root. + + * `gcd_domain_unique` : The minimal polynomial of an element `x` is uniquely characterized by + its defining property: if there is another monic polynomial of minimal degree that has `x` as a + root, then this polynomial is equal to the minimal polynomial of `x`. +-/ + +open_locale classical polynomial +open polynomial set function minpoly + +namespace minpoly + + +section gcd_domain + +variables {R S : Type*} (K L : Type*) [comm_ring R] [is_domain R] [normalized_gcd_monoid R] + [field K] [comm_ring S] [is_domain S] [algebra R K] [is_fraction_ring R K] [algebra R S] [field L] + [algebra S L] [algebra K L] [algebra R L] [is_scalar_tower R K L] [is_scalar_tower R S L] + {s : S} (hs : is_integral R s) + +include hs + +/-- For GCD domains, the minimal polynomial over the ring is the same as the minimal polynomial +over the fraction field. See `minpoly.gcd_domain_eq_field_fractions'` if `S` is already a +`K`-algebra. -/ +lemma gcd_domain_eq_field_fractions : + minpoly K (algebra_map S L s) = (minpoly R s).map (algebra_map R K) := +begin + refine (eq_of_irreducible_of_monic _ _ _).symm, + { exact (polynomial.is_primitive.irreducible_iff_irreducible_map_fraction_map + (polynomial.monic.is_primitive (monic hs))).1 (irreducible hs) }, + { rw [aeval_map_algebra_map, aeval_algebra_map_apply, aeval, map_zero] }, + { exact (monic hs).map _ } +end + +/-- For GCD domains, the minimal polynomial over the ring is the same as the minimal polynomial +over the fraction field. Compared to `minpoly.gcd_domain_eq_field_fractions`, this version is useful +if the element is in a ring that is already a `K`-algebra. -/ +lemma gcd_domain_eq_field_fractions' [algebra K S] [is_scalar_tower R K S] : + minpoly K s = (minpoly R s).map (algebra_map R K) := +begin + let L := fraction_ring S, + rw [← gcd_domain_eq_field_fractions K L hs], + refine minpoly.eq_of_algebra_map_eq (is_fraction_ring.injective S L) + (is_integral_of_is_scalar_tower hs) rfl +end + +variable [no_zero_smul_divisors R S] + +/-- For GCD domains, the minimal polynomial divides any primitive polynomial that has the integral +element as root. See also `minpoly.dvd` which relaxes the assumptions on `S` in exchange for +stronger assumptions on `R`. -/ +lemma gcd_domain_dvd {P : R[X]} (hP : P ≠ 0) (hroot : polynomial.aeval s P = 0) : minpoly R s ∣ P := +begin + let K := fraction_ring R, + let L := fraction_ring S, + let P₁ := P.prim_part, + suffices : minpoly R s ∣ P₁, + { exact dvd_trans this (prim_part_dvd _) }, + apply (is_primitive.dvd_iff_fraction_map_dvd_fraction_map K (monic hs).is_primitive + P.is_primitive_prim_part).2, + let y := algebra_map S L s, + have hy : is_integral R y := hs.algebra_map, + rw [← gcd_domain_eq_field_fractions K L hs], + refine dvd _ _ _, + rw [aeval_map_algebra_map, aeval_algebra_map_apply, aeval_prim_part_eq_zero hP hroot, map_zero] +end + +/-- If an element `x` is a root of a nonzero polynomial `p`, then the degree of `p` is at least the +degree of the minimal polynomial of `x`. See also `minpoly.degree_le_of_ne_zero` which relaxes the +assumptions on `S` in exchange for stronger assumptions on `R`. -/ +lemma gcd_domain_degree_le_of_ne_zero {p : R[X]} (hp0 : p ≠ 0) (hp : polynomial.aeval s p = 0) : + degree (minpoly R s) ≤ degree p := +begin + rw [degree_eq_nat_degree (minpoly.ne_zero hs), degree_eq_nat_degree hp0], + norm_cast, + exact nat_degree_le_of_dvd (gcd_domain_dvd hs hp0 hp) hp0 +end + +omit hs + +/-- The minimal polynomial of an element `x` is uniquely characterized by its defining property: +if there is another monic polynomial of minimal degree that has `x` as a root, then this polynomial +is equal to the minimal polynomial of `x`. See also `minpoly.unique` which relaxes the +assumptions on `S` in exchange for stronger assumptions on `R`. -/ +lemma gcd_domain_unique {P : R[X]} (hmo : P.monic) (hP : polynomial.aeval s P = 0) + (Pmin : ∀ Q : R[X], Q.monic → polynomial.aeval s Q = 0 → degree P ≤ degree Q) : + P = minpoly R s := +begin + have hs : is_integral R s := ⟨P, hmo, hP⟩, + symmetry, apply eq_of_sub_eq_zero, + by_contra hnz, + have := gcd_domain_degree_le_of_ne_zero hs hnz (by simp [hP]), + contrapose! this, + refine degree_sub_lt _ (ne_zero hs) _, + { exact le_antisymm (min R s hmo hP) + (Pmin (minpoly R s) (monic hs) (aeval R s)) }, + { rw [(monic hs).leading_coeff, hmo.leading_coeff] } +end + +end gcd_domain + +end minpoly diff --git a/src/field_theory/separable.lean b/src/field_theory/separable.lean index 3040b17edc7e3..ffbe6f49b9320 100644 --- a/src/field_theory/separable.lean +++ b/src/field_theory/separable.lean @@ -7,7 +7,7 @@ Authors: Kenny Lau import algebra.squarefree import data.polynomial.expand import data.polynomial.splits -import field_theory.minpoly +import field_theory.minpoly.field import ring_theory.power_basis /-! diff --git a/src/linear_algebra/annihilating_polynomial.lean b/src/linear_algebra/annihilating_polynomial.lean index 8a868acb78e22..784fee1a4b31b 100644 --- a/src/linear_algebra/annihilating_polynomial.lean +++ b/src/linear_algebra/annihilating_polynomial.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Justin Thomas. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Justin Thomas -/ -import field_theory.minpoly +import field_theory.minpoly.field import ring_theory.principal_ideal_domain /-! diff --git a/src/linear_algebra/charpoly/basic.lean b/src/linear_algebra/charpoly/basic.lean index fe2808b923912..8941aab44f79e 100644 --- a/src/linear_algebra/charpoly/basic.lean +++ b/src/linear_algebra/charpoly/basic.lean @@ -6,7 +6,7 @@ Authors: Riccardo Brasca import linear_algebra.free_module.finite.basic import linear_algebra.matrix.charpoly.coeff -import field_theory.minpoly +import field_theory.minpoly.field /-! diff --git a/src/ring_theory/adjoin_root.lean b/src/ring_theory/adjoin_root.lean index a3db61356c49e..ae14a401494fd 100644 --- a/src/ring_theory/adjoin_root.lean +++ b/src/ring_theory/adjoin_root.lean @@ -5,6 +5,7 @@ Authors: Mario Carneiro, Chris Hughes -/ import algebra.algebra.basic import data.polynomial.field_division +import field_theory.minpoly.gcd_monoid import ring_theory.adjoin.basic import ring_theory.finite_presentation import ring_theory.finite_type diff --git a/src/ring_theory/power_basis.lean b/src/ring_theory/power_basis.lean index 349cf5064473d..fb053509a8d99 100644 --- a/src/ring_theory/power_basis.lean +++ b/src/ring_theory/power_basis.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 Anne Baanen. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Anne Baanen -/ -import field_theory.minpoly +import field_theory.minpoly.field /-! # Power basis From 6d0adfa76594f304b4650d098273d4366edeb61b Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Wed, 4 Jan 2023 14:42:24 +0000 Subject: [PATCH 0140/1291] feat(geometry/euclidean/angle/sphere): law of sines, converse of inscribed angles (#17992) Add some (oriented angle) versions of the law of sines / sine rule, and of the converse of angles in the same segment / opposite angles of a cyclic quadrilateral. As with previous lemmas in this area, more versions would be appropriate to add in future, including unoriented angle versions and versions of the law of sines that don't involve the circumradius and so can be applied even in the degenerate case of collinear points, but this seems a reasonable starting point. --- src/geometry/euclidean/angle/sphere.lean | 321 ++++++++++++++++++++++- 1 file changed, 317 insertions(+), 4 deletions(-) diff --git a/src/geometry/euclidean/angle/sphere.lean b/src/geometry/euclidean/angle/sphere.lean index 27ffa53c11acd..9293c24e9f52a 100644 --- a/src/geometry/euclidean/angle/sphere.lean +++ b/src/geometry/euclidean/angle/sphere.lean @@ -3,8 +3,8 @@ Copyright (c) 2022 Joseph Myers. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Joseph Myers -/ -import geometry.euclidean.angle.oriented.affine -import geometry.euclidean.basic +import geometry.euclidean.angle.oriented.right_angle +import geometry.euclidean.circumcenter /-! # Angles in circles and sphere. @@ -73,8 +73,10 @@ include hd2 local notation `o` := module.oriented.positive_orientation +namespace sphere + /-- Angle at center of a circle equals twice angle at circumference, oriented angle version. -/ -lemma sphere.oangle_center_eq_two_zsmul_oangle {s : sphere P} {p₁ p₂ p₃ : P} (hp₁ : p₁ ∈ s) +lemma oangle_center_eq_two_zsmul_oangle {s : sphere P} {p₁ p₂ p₃ : P} (hp₁ : p₁ ∈ s) (hp₂ : p₂ ∈ s) (hp₃ : p₃ ∈ s) (hp₂p₁ : p₂ ≠ p₁) (hp₂p₃ : p₂ ≠ p₃) : ∡ p₁ s.center p₃ = (2 : ℤ) • ∡ p₁ p₂ p₃ := begin @@ -86,7 +88,7 @@ end /-- Oriented angle version of "angles in same segment are equal" and "opposite angles of a cyclic quadrilateral add to π", for oriented angles mod π (for which those are the same result), represented here as equality of twice the angles. -/ -lemma sphere.two_zsmul_oangle_eq {s : sphere P} {p₁ p₂ p₃ p₄ : P} (hp₁ : p₁ ∈ s) (hp₂ : p₂ ∈ s) +lemma two_zsmul_oangle_eq {s : sphere P} {p₁ p₂ p₃ p₄ : P} (hp₁ : p₁ ∈ s) (hp₂ : p₂ ∈ s) (hp₃ : p₃ ∈ s) (hp₄ : p₄ ∈ s) (hp₂p₁ : p₂ ≠ p₁) (hp₂p₄ : p₂ ≠ p₄) (hp₃p₁ : p₃ ≠ p₁) (hp₃p₄ : p₃ ≠ p₄) : (2 : ℤ) • ∡ p₁ p₂ p₄ = (2 : ℤ) • ∡ p₁ p₃ p₄ := begin @@ -97,6 +99,8 @@ begin simp [hp₂p₁, hp₂p₄, hp₃p₁, hp₃p₄] end +end sphere + /-- Oriented angle version of "angles in same segment are equal" and "opposite angles of a cyclic quadrilateral add to π", for oriented angles mod π (for which those are the same result), represented here as equality of twice the angles. -/ @@ -109,4 +113,313 @@ begin exact sphere.two_zsmul_oangle_eq hs.1 hs.2.1 hs.2.2.1 hs.2.2.2 hp₂p₁ hp₂p₄ hp₃p₁ hp₃p₄ end +namespace sphere + +/-- The angle at the apex of an isosceles triangle is `π` minus twice a base angle, oriented +angle-at-point form where the apex is given as the center of a circle. -/ +lemma oangle_eq_pi_sub_two_zsmul_oangle_center_left {s : sphere P} {p₁ p₂ : P} (hp₁ : p₁ ∈ s) + (hp₂ : p₂ ∈ s) (h : p₁ ≠ p₂) : ∡ p₁ s.center p₂ = π - (2 : ℤ) • ∡ s.center p₂ p₁ := +by rw [oangle_eq_pi_sub_two_zsmul_oangle_of_dist_eq h.symm + (dist_center_eq_dist_center_of_mem_sphere' hp₂ hp₁)] + +/-- The angle at the apex of an isosceles triangle is `π` minus twice a base angle, oriented +angle-at-point form where the apex is given as the center of a circle. -/ +lemma oangle_eq_pi_sub_two_zsmul_oangle_center_right {s : sphere P} {p₁ p₂ : P} (hp₁ : p₁ ∈ s) + (hp₂ : p₂ ∈ s) (h : p₁ ≠ p₂) : ∡ p₁ s.center p₂ = π - (2 : ℤ) • ∡ p₂ p₁ s.center := +by rw [oangle_eq_pi_sub_two_zsmul_oangle_center_left hp₁ hp₂ h, + oangle_eq_oangle_of_dist_eq (dist_center_eq_dist_center_of_mem_sphere' hp₂ hp₁)] + +/-- Twice a base angle of an isosceles triangle with apex at the center of a circle, plus twice +the angle at the apex of a triangle with the same base but apex on the circle, equals `π`. -/ +lemma two_zsmul_oangle_center_add_two_zsmul_oangle_eq_pi {s : sphere P} {p₁ p₂ p₃ : P} + (hp₁ : p₁ ∈ s) (hp₂ : p₂ ∈ s) (hp₃ : p₃ ∈ s) (hp₂p₁ : p₂ ≠ p₁) (hp₂p₃ : p₂ ≠ p₃) + (hp₁p₃ : p₁ ≠ p₃) : (2 : ℤ) • ∡ p₃ p₁ s.center + (2 : ℤ) • ∡ p₁ p₂ p₃ = π := +by rw [←oangle_center_eq_two_zsmul_oangle hp₁ hp₂ hp₃ hp₂p₁ hp₂p₃, + oangle_eq_pi_sub_two_zsmul_oangle_center_right hp₁ hp₃ hp₁p₃, add_sub_cancel'_right] + +/-- A base angle of an isosceles triangle with apex at the center of a circle is acute. -/ +lemma abs_oangle_center_left_to_real_lt_pi_div_two {s : sphere P} {p₁ p₂ : P} (hp₁ : p₁ ∈ s) + (hp₂ : p₂ ∈ s) : |(∡ s.center p₂ p₁).to_real| < π / 2 := +abs_oangle_right_to_real_lt_pi_div_two_of_dist_eq + (dist_center_eq_dist_center_of_mem_sphere' hp₂ hp₁) + +/-- A base angle of an isosceles triangle with apex at the center of a circle is acute. -/ +lemma abs_oangle_center_right_to_real_lt_pi_div_two {s : sphere P} {p₁ p₂ : P} (hp₁ : p₁ ∈ s) + (hp₂ : p₂ ∈ s) : |(∡ p₂ p₁ s.center).to_real| < π / 2 := +abs_oangle_left_to_real_lt_pi_div_two_of_dist_eq + (dist_center_eq_dist_center_of_mem_sphere' hp₂ hp₁) + +/-- Given two points on a circle, the center of that circle may be expressed explicitly as a +multiple (by half the tangent of the angle between the chord and the radius at one of those +points) of a `π / 2` rotation of the vector between those points, plus the midpoint of those +points. -/ +lemma tan_div_two_smul_rotation_pi_div_two_vadd_midpoint_eq_center {s : sphere P} {p₁ p₂ : P} + (hp₁ : p₁ ∈ s) (hp₂ : p₂ ∈ s) (h : p₁ ≠ p₂) : + (real.angle.tan (∡ p₂ p₁ s.center) / 2) • ((o).rotation (π / 2 : ℝ) (p₂ -ᵥ p₁)) +ᵥ + midpoint ℝ p₁ p₂ = s.center := +begin + obtain ⟨r, hr⟩ := (dist_eq_iff_eq_smul_rotation_pi_div_two_vadd_midpoint h).1 + (dist_center_eq_dist_center_of_mem_sphere hp₁ hp₂), + rw [←hr, ←oangle_midpoint_rev_left, oangle, vadd_vsub_assoc], + nth_rewrite 0 (show p₂ -ᵥ p₁ = (2 : ℝ) • (midpoint ℝ p₁ p₂ -ᵥ p₁), by simp), + rw [map_smul, smul_smul, add_comm, (o).tan_oangle_add_right_smul_rotation_pi_div_two, + mul_div_cancel _ (two_ne_zero' ℝ)], + simpa using h.symm +end + +/-- Given three points on a circle, the center of that circle may be expressed explicitly as a +multiple (by half the inverse of the tangent of the angle at one of those points) of a `π / 2` +rotation of the vector between the other two points, plus the midpoint of those points. -/ +lemma inv_tan_div_two_smul_rotation_pi_div_two_vadd_midpoint_eq_center {s : sphere P} + {p₁ p₂ p₃ : P} (hp₁ : p₁ ∈ s) (hp₂ : p₂ ∈ s) (hp₃ : p₃ ∈ s) (hp₁p₂ : p₁ ≠ p₂) (hp₁p₃ : p₁ ≠ p₃) + (hp₂p₃ : p₂ ≠ p₃) : + ((real.angle.tan (∡ p₁ p₂ p₃))⁻¹ / 2) • ((o).rotation (π / 2 : ℝ) (p₃ -ᵥ p₁)) +ᵥ + midpoint ℝ p₁ p₃ = s.center := +begin + convert tan_div_two_smul_rotation_pi_div_two_vadd_midpoint_eq_center hp₁ hp₃ hp₁p₃, + convert (real.angle.tan_eq_inv_of_two_zsmul_add_two_zsmul_eq_pi _).symm, + rw [add_comm, + two_zsmul_oangle_center_add_two_zsmul_oangle_eq_pi hp₁ hp₂ hp₃ hp₁p₂.symm hp₂p₃ hp₁p₃] +end + +/-- Given two points on a circle, the radius of that circle may be expressed explicitly as half +the distance between those two points divided by the cosine of the angle between the chord and +the radius at one of those points. -/ +lemma dist_div_cos_oangle_center_div_two_eq_radius {s : sphere P} {p₁ p₂ : P} (hp₁ : p₁ ∈ s) + (hp₂ : p₂ ∈ s) (h : p₁ ≠ p₂) : dist p₁ p₂ / real.angle.cos (∡ p₂ p₁ s.center) / 2 = s.radius := +begin + rw [div_right_comm, div_eq_mul_inv _ (2 : ℝ), mul_comm, + (show (2 : ℝ)⁻¹ * dist p₁ p₂ = dist p₁ (midpoint ℝ p₁ p₂), by simp), ←mem_sphere.1 hp₁, + ←tan_div_two_smul_rotation_pi_div_two_vadd_midpoint_eq_center hp₁ hp₂ h, + ←oangle_midpoint_rev_left, oangle, vadd_vsub_assoc, + (show p₂ -ᵥ p₁ = (2 : ℝ) • (midpoint ℝ p₁ p₂ -ᵥ p₁), by simp), map_smul, smul_smul, + div_mul_cancel _ (two_ne_zero' ℝ), @dist_eq_norm_vsub' V, @dist_eq_norm_vsub' V, + vadd_vsub_assoc, add_comm, (o).oangle_add_right_smul_rotation_pi_div_two, + real.angle.cos_coe, real.cos_arctan, one_div, div_inv_eq_mul, + ←mul_self_inj (mul_nonneg (norm_nonneg _) (real.sqrt_nonneg _)) (norm_nonneg _), + norm_add_sq_eq_norm_sq_add_norm_sq_real ((o).inner_smul_rotation_pi_div_two_right _ _), + ←mul_assoc, mul_comm, mul_comm _ (real.sqrt _), ←mul_assoc, ←mul_assoc, + real.mul_self_sqrt (add_nonneg zero_le_one (sq_nonneg _)), norm_smul, + linear_isometry_equiv.norm_map], + swap, { simpa using h.symm }, + conv_rhs { rw [←mul_assoc, mul_comm _ ‖real.angle.tan _‖, ←mul_assoc, real.norm_eq_abs, + abs_mul_abs_self] }, + ring +end + +/-- Given two points on a circle, twice the radius of that circle may be expressed explicitly as +the distance between those two points divided by the cosine of the angle between the chord and +the radius at one of those points. -/ +lemma dist_div_cos_oangle_center_eq_two_mul_radius {s : sphere P} {p₁ p₂ : P} (hp₁ : p₁ ∈ s) + (hp₂ : p₂ ∈ s) (h : p₁ ≠ p₂) : dist p₁ p₂ / real.angle.cos (∡ p₂ p₁ s.center) = 2 * s.radius := +by rw [←dist_div_cos_oangle_center_div_two_eq_radius hp₁ hp₂ h, + mul_div_cancel' _ (two_ne_zero' ℝ)] + +/-- Given three points on a circle, the radius of that circle may be expressed explicitly as half +the distance between two of those points divided by the absolute value of the sine of the angle +at the third point (a version of the law of sines or sine rule). -/ +lemma dist_div_sin_oangle_div_two_eq_radius {s : sphere P} {p₁ p₂ p₃ : P} (hp₁ : p₁ ∈ s) + (hp₂ : p₂ ∈ s) (hp₃ : p₃ ∈ s) (hp₁p₂ : p₁ ≠ p₂) (hp₁p₃ : p₁ ≠ p₃) (hp₂p₃ : p₂ ≠ p₃) : + dist p₁ p₃ / |real.angle.sin (∡ p₁ p₂ p₃)| / 2 = s.radius := +begin + convert dist_div_cos_oangle_center_div_two_eq_radius hp₁ hp₃ hp₁p₃, + rw [←real.angle.abs_cos_eq_abs_sin_of_two_zsmul_add_two_zsmul_eq_pi + (two_zsmul_oangle_center_add_two_zsmul_oangle_eq_pi + hp₁ hp₂ hp₃ hp₁p₂.symm hp₂p₃ hp₁p₃), + _root_.abs_of_nonneg (real.angle.cos_nonneg_iff_abs_to_real_le_pi_div_two.2 _)], + exact (abs_oangle_center_right_to_real_lt_pi_div_two hp₁ hp₃).le +end + +/-- Given three points on a circle, twice the radius of that circle may be expressed explicitly as +the distance between two of those points divided by the absolute value of the sine of the angle +at the third point (a version of the law of sines or sine rule). -/ +lemma dist_div_sin_oangle_eq_two_mul_radius {s : sphere P} {p₁ p₂ p₃ : P} (hp₁ : p₁ ∈ s) + (hp₂ : p₂ ∈ s) (hp₃ : p₃ ∈ s) (hp₁p₂ : p₁ ≠ p₂) (hp₁p₃ : p₁ ≠ p₃) (hp₂p₃ : p₂ ≠ p₃) : + dist p₁ p₃ / |real.angle.sin (∡ p₁ p₂ p₃)| = 2 * s.radius := +by rw [←dist_div_sin_oangle_div_two_eq_radius hp₁ hp₂ hp₃ hp₁p₂ hp₁p₃ hp₂p₃, + mul_div_cancel' _ (two_ne_zero' ℝ)] + +end sphere + +end euclidean_geometry + +namespace affine +namespace triangle + +open euclidean_geometry + +variables {V : Type*} {P : Type*} [inner_product_space ℝ V] [metric_space P] +variables [normed_add_torsor V P] [hd2 : fact (finrank ℝ V = 2)] [module.oriented ℝ V (fin 2)] +include hd2 + +local notation `o` := module.oriented.positive_orientation + +/-- The circumcenter of a triangle may be expressed explicitly as a multiple (by half the inverse +of the tangent of the angle at one of the vertices) of a `π / 2` rotation of the vector between +the other two vertices, plus the midpoint of those vertices. -/ +lemma inv_tan_div_two_smul_rotation_pi_div_two_vadd_midpoint_eq_circumcenter (t : triangle ℝ P) + {i₁ i₂ i₃ : fin 3} (h₁₂ : i₁ ≠ i₂) (h₁₃ : i₁ ≠ i₃) (h₂₃ : i₂ ≠ i₃) : + ((real.angle.tan (∡ (t.points i₁) (t.points i₂) (t.points i₃)))⁻¹ / 2) • + ((o).rotation (π / 2 : ℝ) (t.points i₃ -ᵥ t.points i₁)) +ᵥ + midpoint ℝ (t.points i₁) (t.points i₃) = t.circumcenter := +sphere.inv_tan_div_two_smul_rotation_pi_div_two_vadd_midpoint_eq_center + (t.mem_circumsphere _) (t.mem_circumsphere _) (t.mem_circumsphere _) + (t.independent.injective.ne h₁₂) (t.independent.injective.ne h₁₃) (t.independent.injective.ne h₂₃) + +/-- The circumradius of a triangle may be expressed explicitly as half the length of a side +divided by the absolute value of the sine of the angle at the third point (a version of the law +of sines or sine rule). -/ +lemma dist_div_sin_oangle_div_two_eq_circumradius (t : triangle ℝ P) {i₁ i₂ i₃ : fin 3} + (h₁₂ : i₁ ≠ i₂) (h₁₃ : i₁ ≠ i₃) (h₂₃ : i₂ ≠ i₃) : + dist (t.points i₁) (t.points i₃) / + |real.angle.sin (∡ (t.points i₁) (t.points i₂) (t.points i₃))| / 2 = t.circumradius := +sphere.dist_div_sin_oangle_div_two_eq_radius (t.mem_circumsphere _) (t.mem_circumsphere _) + (t.mem_circumsphere _) (t.independent.injective.ne h₁₂) (t.independent.injective.ne h₁₃) + (t.independent.injective.ne h₂₃) + +/-- Twice the circumradius of a triangle may be expressed explicitly as the length of a side +divided by the absolute value of the sine of the angle at the third point (a version of the law +of sines or sine rule). -/ +lemma dist_div_sin_oangle_eq_two_mul_circumradius (t : triangle ℝ P) {i₁ i₂ i₃ : fin 3} + (h₁₂ : i₁ ≠ i₂) (h₁₃ : i₁ ≠ i₃) (h₂₃ : i₂ ≠ i₃) : + dist (t.points i₁) (t.points i₃) / + |real.angle.sin (∡ (t.points i₁) (t.points i₂) (t.points i₃))| = 2 * t.circumradius := +sphere.dist_div_sin_oangle_eq_two_mul_radius (t.mem_circumsphere _) (t.mem_circumsphere _) + (t.mem_circumsphere _) (t.independent.injective.ne h₁₂) (t.independent.injective.ne h₁₃) + (t.independent.injective.ne h₂₃) + +/-- The circumsphere of a triangle may be expressed explicitly in terms of two points and the +angle at the third point. -/ +lemma circumsphere_eq_of_dist_of_oangle (t : triangle ℝ P) {i₁ i₂ i₃ : fin 3} (h₁₂ : i₁ ≠ i₂) + (h₁₃ : i₁ ≠ i₃) (h₂₃ : i₂ ≠ i₃) : + t.circumsphere = ⟨((real.angle.tan (∡ (t.points i₁) (t.points i₂) (t.points i₃)))⁻¹ / 2) • + ((o).rotation (π / 2 : ℝ) (t.points i₃ -ᵥ t.points i₁)) +ᵥ + midpoint ℝ (t.points i₁) (t.points i₃), + dist (t.points i₁) (t.points i₃) / + |real.angle.sin (∡ (t.points i₁) (t.points i₂) (t.points i₃))| / 2⟩ := +t.circumsphere.ext _ + (t.inv_tan_div_two_smul_rotation_pi_div_two_vadd_midpoint_eq_circumcenter h₁₂ h₁₃ h₂₃).symm + (t.dist_div_sin_oangle_div_two_eq_circumradius h₁₂ h₁₃ h₂₃).symm + +/-- If two triangles have two points the same, and twice the angle at the third point the same, +they have the same circumsphere. -/ +lemma circumsphere_eq_circumsphere_of_eq_of_eq_of_two_zsmul_oangle_eq {t₁ t₂ : triangle ℝ P} + {i₁ i₂ i₃ : fin 3} (h₁₂ : i₁ ≠ i₂) (h₁₃ : i₁ ≠ i₃) (h₂₃ : i₂ ≠ i₃) + (h₁ : t₁.points i₁ = t₂.points i₁) (h₃ : t₁.points i₃ = t₂.points i₃) + (h₂ : (2 : ℤ) • ∡ (t₁.points i₁) (t₁.points i₂) (t₁.points i₃) = + (2 : ℤ) • ∡ (t₂.points i₁) (t₂.points i₂) (t₂.points i₃)) : + t₁.circumsphere = t₂.circumsphere := +begin + rw [t₁.circumsphere_eq_of_dist_of_oangle h₁₂ h₁₃ h₂₃, + t₂.circumsphere_eq_of_dist_of_oangle h₁₂ h₁₃ h₂₃], + congrm ⟨((_ : ℝ)⁻¹ / 2) • _ +ᵥ _, _ / _ / 2⟩, + { exact real.angle.tan_eq_of_two_zsmul_eq h₂ }, + { rw [h₁, h₃] }, + { rw [h₁, h₃] }, + { rw [h₁, h₃] }, + { exact real.angle.abs_sin_eq_of_two_zsmul_eq h₂ } +end + +/-- Given a triangle, and a fourth point such that twice the angle between two points of the +triangle at that fourth point equals twice the third angle of the triangle, the fourth point +lies in the circumsphere of the triangle. -/ +lemma mem_circumsphere_of_two_zsmul_oangle_eq {t : triangle ℝ P} {p : P} {i₁ i₂ i₃ : fin 3} + (h₁₂ : i₁ ≠ i₂) (h₁₃ : i₁ ≠ i₃) (h₂₃ : i₂ ≠ i₃) + (h : (2 : ℤ) • ∡ (t.points i₁) p (t.points i₃) = + (2 : ℤ) • ∡ (t.points i₁) (t.points i₂) (t.points i₃)) : p ∈ t.circumsphere := +begin + let t'p : fin 3 → P := function.update t.points i₂ p, + have h₁ : t'p i₁ = t.points i₁, { simp [t'p, h₁₂] }, + have h₂ : t'p i₂ = p, { simp [t'p] }, + have h₃ : t'p i₃ = t.points i₃, { simp [t'p, h₂₃.symm] }, + have ha : affine_independent ℝ t'p, + { rw [affine_independent_iff_not_collinear_of_ne h₁₂ h₁₃ h₂₃, h₁, h₂, h₃, + collinear_iff_of_two_zsmul_oangle_eq h, + ←affine_independent_iff_not_collinear_of_ne h₁₂ h₁₃ h₂₃], + exact t.independent }, + let t' : triangle ℝ P := ⟨t'p, ha⟩, + have h₁' : t'.points i₁ = t.points i₁ := h₁, + have h₂' : t'.points i₂ = p := h₂, + have h₃' : t'.points i₃ = t.points i₃ := h₃, + have h' : (2 : ℤ) • ∡ (t'.points i₁) (t'.points i₂) (t'.points i₃) = + (2 : ℤ) • ∡ (t.points i₁) (t.points i₂) (t.points i₃), { rwa [h₁', h₂', h₃'] }, + rw [←circumsphere_eq_circumsphere_of_eq_of_eq_of_two_zsmul_oangle_eq h₁₂ h₁₃ h₂₃ h₁' h₃' h', + ←h₂'], + exact simplex.mem_circumsphere _ _ +end + +end triangle +end affine + +namespace euclidean_geometry + +variables {V : Type*} {P : Type*} [inner_product_space ℝ V] [metric_space P] +variables [normed_add_torsor V P] [hd2 : fact (finrank ℝ V = 2)] [module.oriented ℝ V (fin 2)] +include hd2 + +local notation `o` := module.oriented.positive_orientation + +/-- Converse of "angles in same segment are equal" and "opposite angles of a cyclic quadrilateral +add to π", for oriented angles mod π. -/ +lemma cospherical_of_two_zsmul_oangle_eq_of_not_collinear {p₁ p₂ p₃ p₄ : P} + (h : (2 : ℤ) • ∡ p₁ p₂ p₄ = (2 : ℤ) • ∡ p₁ p₃ p₄) (hn : ¬collinear ℝ ({p₁, p₂, p₄} : set P)) : + cospherical ({p₁, p₂, p₃, p₄} : set P) := +begin + have hn' : ¬collinear ℝ ({p₁, p₃, p₄} : set P), { rwa ←collinear_iff_of_two_zsmul_oangle_eq h }, + let t₁ : affine.triangle ℝ P := ⟨![p₁, p₂, p₄], affine_independent_iff_not_collinear_set.2 hn⟩, + let t₂ : affine.triangle ℝ P := ⟨![p₁, p₃, p₄], affine_independent_iff_not_collinear_set.2 hn'⟩, + rw cospherical_iff_exists_sphere, + refine ⟨t₂.circumsphere, _⟩, + simp_rw [set.insert_subset, set.singleton_subset_iff], + refine ⟨t₂.mem_circumsphere 0, _, t₂.mem_circumsphere 1, t₂.mem_circumsphere 2⟩, + rw affine.triangle.circumsphere_eq_circumsphere_of_eq_of_eq_of_two_zsmul_oangle_eq + (dec_trivial : (0 : fin 3) ≠ 1) (dec_trivial: (0 : fin 3) ≠ 2) dec_trivial + (show t₂.points 0 = t₁.points 0, from rfl) rfl h.symm, + exact t₁.mem_circumsphere 1 +end + +/-- Converse of "angles in same segment are equal" and "opposite angles of a cyclic quadrilateral +add to π", for oriented angles mod π, with a "concyclic" conclusion. -/ +lemma concyclic_of_two_zsmul_oangle_eq_of_not_collinear {p₁ p₂ p₃ p₄ : P} + (h : (2 : ℤ) • ∡ p₁ p₂ p₄ = (2 : ℤ) • ∡ p₁ p₃ p₄) (hn : ¬collinear ℝ ({p₁, p₂, p₄} : set P)) : + concyclic ({p₁, p₂, p₃, p₄} : set P) := +⟨cospherical_of_two_zsmul_oangle_eq_of_not_collinear h hn, coplanar_of_fact_finrank_eq_two _⟩ + +/-- Converse of "angles in same segment are equal" and "opposite angles of a cyclic quadrilateral +add to π", for oriented angles mod π, with a "cospherical or collinear" conclusion. -/ +lemma cospherical_or_collinear_of_two_zsmul_oangle_eq {p₁ p₂ p₃ p₄ : P} + (h : (2 : ℤ) • ∡ p₁ p₂ p₄ = (2 : ℤ) • ∡ p₁ p₃ p₄) : + cospherical ({p₁, p₂, p₃, p₄} : set P) ∨ collinear ℝ ({p₁, p₂, p₃, p₄} : set P) := +begin + by_cases hc : collinear ℝ ({p₁, p₂, p₄} : set P), + { by_cases he : p₁ = p₄, + { rw [he, set.insert_eq_self.2 (set.mem_insert_of_mem _ (set.mem_insert_of_mem _ + (set.mem_singleton _)))], + by_cases hl : collinear ℝ ({p₂, p₃, p₄} : set P), { exact or.inr hl }, + rw or_iff_left hl, + let t : affine.triangle ℝ P := ⟨![p₂, p₃, p₄], affine_independent_iff_not_collinear_set.2 hl⟩, + rw cospherical_iff_exists_sphere, + refine ⟨t.circumsphere, _⟩, + simp_rw [set.insert_subset, set.singleton_subset_iff], + exact ⟨t.mem_circumsphere 0, t.mem_circumsphere 1, t.mem_circumsphere 2⟩ }, + have hc' : collinear ℝ ({p₁, p₃, p₄} : set P), + { rwa [←collinear_iff_of_two_zsmul_oangle_eq h] }, + refine or.inr _, + rw set.insert_comm p₁ p₂ at hc, + rwa [set.insert_comm p₁ p₂, hc'.collinear_insert_iff_of_ne (set.mem_insert _ _) + (set.mem_insert_of_mem _ (set.mem_insert_of_mem _ (set.mem_singleton _))) he] }, + { exact or.inl (cospherical_of_two_zsmul_oangle_eq_of_not_collinear h hc) } +end + +/-- Converse of "angles in same segment are equal" and "opposite angles of a cyclic quadrilateral +add to π", for oriented angles mod π, with a "concyclic or collinear" conclusion. -/ +lemma concyclic_or_collinear_of_two_zsmul_oangle_eq {p₁ p₂ p₃ p₄ : P} + (h : (2 : ℤ) • ∡ p₁ p₂ p₄ = (2 : ℤ) • ∡ p₁ p₃ p₄) : + concyclic ({p₁, p₂, p₃, p₄} : set P) ∨ collinear ℝ ({p₁, p₂, p₃, p₄} : set P) := +begin + rcases cospherical_or_collinear_of_two_zsmul_oangle_eq h with hc | hc, + { exact or.inl ⟨hc, coplanar_of_fact_finrank_eq_two _⟩ }, + { exact or.inr hc } +end + end euclidean_geometry From d1723c047a091ae3fca0af8aeab1743e1a898611 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 4 Jan 2023 22:44:02 +0000 Subject: [PATCH 0141/1291] refactor(group_theory/perm/cycle/basic): Consolidate API (#17898) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Reorganise and complete the cycle API: * Previously, `is_cycle` was spelling out the definition of `same_cycle`. Now use it explicitly. * Change binder to semi-implicit in the definition of `is_cycle`. * Add lemmas and iff aliases. * Golf existing proofs using those (mostly invisible from the diff because git decided I am moving the `is_cycle` API) * Improve lemma names, mostly for better dot notation. ## New lemmas * `maps_to.extend_domain` * `surj_on.extend_domain` * `bij_on.extend_domain` * `extend_domain_pow` * `extend_domain_zpow` * `same_cycle.rfl` * `eq.same_cycle` * `same_cycle.conj` * `same_cycle_conj` * `same_cycle_pow_right_iff` * `same_cycle_zpow_right_iff` * `same_cycle.pow_left` * `same_cycle.pow_right` * `same_cycle.zpow_left` * `same_cycle.zpow_left` * `same_cycle.of_pow` * `same_cycle.of_zpow` * `same_cycle_subtype_perm` * `same_cycle.subtype_perm` * `same_cycle_extend_domain` * `is_cycle.conj` * `is_cycle.pow_eq_one_iff'` * `is_cycle.pow_eq_one_iff''` ## Renames * `order_of_is_cycle` → `is_cycle.order_of` * `is_cycle.is_cycle_conj` → `is_cycle.conj` * `is_cycle_of_is_cycle_pow` → `is_cycle.of_pow` * `is_cycle_of_is_cycle_zpow` → `is_cycle.of_zpow` * `same_cycle_cycle` → `is_cycle_iff_same_cycle` * `mem_support_pos_pow_iff_of_lt_order_of` → `support_pow_of_pos_of_lt_order_of` and change the statement to talk about unextensionalised finset equality --- src/data/set/function.lean | 25 +- src/group_theory/perm/basic.lean | 6 + src/group_theory/perm/cycle/basic.lean | 459 ++++++++++++---------- src/group_theory/perm/cycle/concrete.lean | 8 +- src/group_theory/perm/cycle/type.lean | 8 +- 5 files changed, 279 insertions(+), 227 deletions(-) diff --git a/src/data/set/function.lean b/src/data/set/function.lean index a2e5964ac4191..c23e7bd94023b 100644 --- a/src/data/set/function.lean +++ b/src/data/set/function.lean @@ -37,7 +37,7 @@ universes u v w x y variables {α : Type u} {β : Type v} {π : α → Type v} {γ : Type w} {ι : Sort x} -open function +open equiv equiv.perm function namespace set @@ -1307,3 +1307,26 @@ lemma antitone_on_of_right_inv_on_of_maps_to (monotone_on_of_right_inv_on_of_maps_to hφ.dual_left φψs ψts).dual_right end function + +/-! ### Equivalences, permutations -/ + +namespace set +variables {p : β → Prop} [decidable_pred p] {f : α ≃ subtype p} {g : perm α} {s t : set α} + +protected lemma maps_to.extend_domain (h : maps_to g s t) : + maps_to (g.extend_domain f) (coe ∘ f '' s) (coe ∘ f '' t) := +by { rintro _ ⟨a, ha, rfl⟩, exact ⟨_, h ha, by rw extend_domain_apply_image⟩ } + +protected lemma surj_on.extend_domain (h : surj_on g s t) : + surj_on (g.extend_domain f) (coe ∘ f '' s) (coe ∘ f '' t) := +begin + rintro _ ⟨a, ha, rfl⟩, + obtain ⟨b, hb, rfl⟩ := h ha, + exact ⟨_, ⟨_, hb, rfl⟩, by rw extend_domain_apply_image⟩, +end + +protected lemma bij_on.extend_domain (h : set.bij_on g s t) : + bij_on (g.extend_domain f) (coe ∘ f '' s) (coe ∘ f '' t) := +⟨h.maps_to.extend_domain, (g.extend_domain f).injective.inj_on _, h.surj_on.extend_domain⟩ + +end set diff --git a/src/group_theory/perm/basic.lean b/src/group_theory/perm/basic.lean index 67dae3c1b4efd..844bbea80a7ad 100644 --- a/src/group_theory/perm/basic.lean +++ b/src/group_theory/perm/basic.lean @@ -237,6 +237,12 @@ lemma extend_domain_hom_injective : function.injective (extend_domain_hom f) := e.extend_domain f = 1 ↔ e = 1 := (injective_iff_map_eq_one' (extend_domain_hom f)).mp (extend_domain_hom_injective f) e +@[simp] lemma extend_domain_pow (n : ℕ) : (e ^ n).extend_domain f = e.extend_domain f ^ n := +map_pow (extend_domain_hom f) _ _ + +@[simp] lemma extend_domain_zpow (n : ℤ) : (e ^ n).extend_domain f = e.extend_domain f ^ n := +map_zpow (extend_domain_hom f) _ _ + end extend_domain section subtype diff --git a/src/group_theory/perm/cycle/basic.lean b/src/group_theory/perm/cycle/basic.lean index fd0239172c1b5..c0860c622aa6d 100644 --- a/src/group_theory/perm/cycle/basic.lean +++ b/src/group_theory/perm/cycle/basic.lean @@ -5,6 +5,7 @@ Authors: Chris Hughes -/ import data.finset.noncomm_prod import data.fintype.perm +import data.int.modeq import group_theory.perm.sign import logic.equiv.fintype /-! @@ -37,58 +38,103 @@ The following two definitions require that `β` is a `fintype`: -/ -namespace equiv.perm open equiv function finset variables {α β : Type*} +namespace equiv.perm + /-! ### `same_cycle` -/ section same_cycle variables {f g : perm α} {p : α → Prop} {x y z : α} /-- The equivalence relation indicating that two points are in the same cycle of a permutation. -/ -def same_cycle (f : perm β) (x y : β) : Prop := ∃ i : ℤ, (f ^ i) x = y +def same_cycle (f : perm α) (x y : α) : Prop := ∃ i : ℤ, (f ^ i) x = y -@[refl] lemma same_cycle.refl (f : perm β) (x : β) : same_cycle f x x := ⟨0, rfl⟩ +@[refl] lemma same_cycle.refl (f : perm α) (x : α) : same_cycle f x x := ⟨0, rfl⟩ +lemma same_cycle.rfl : same_cycle f x x := same_cycle.refl _ _ + +protected lemma _root_.eq.same_cycle (h : x = y) (f : perm α) : f.same_cycle x y := by rw h @[symm] lemma same_cycle.symm : same_cycle f x y → same_cycle f y x := λ ⟨i, hi⟩, ⟨-i, by rw [zpow_neg, ← hi, inv_apply_self]⟩ +lemma same_cycle_comm : same_cycle f x y ↔ same_cycle f y x := ⟨same_cycle.symm, same_cycle.symm⟩ + @[trans] lemma same_cycle.trans : same_cycle f x y → same_cycle f y z → same_cycle f x z := λ ⟨i, hi⟩ ⟨j, hj⟩, ⟨j + i, by rw [zpow_add, mul_apply, hi, hj]⟩ -lemma same_cycle_inv (f : perm α) : same_cycle f⁻¹ x y ↔ same_cycle f x y := +@[simp] lemma same_cycle_one : same_cycle 1 x y ↔ x = y := by simp [same_cycle] + +@[simp] lemma same_cycle_inv : same_cycle f⁻¹ x y ↔ same_cycle f x y := (equiv.neg _).exists_congr_left.trans $ by simp [same_cycle] +alias same_cycle_inv ↔ same_cycle.of_inv same_cycle.inv + +@[simp] lemma same_cycle_conj : same_cycle (g * f * g⁻¹) x y ↔ same_cycle f (g⁻¹ x) (g⁻¹ y) := +exists_congr $ λ i, by simp [conj_zpow, eq_inv_iff_eq] + +lemma same_cycle.conj : same_cycle f x y → same_cycle (g * f * g⁻¹) (g x) (g y) := +by simp [same_cycle_conj] + lemma same_cycle.apply_eq_self_iff : same_cycle f x y → (f x = x ↔ f y = y) := λ ⟨i, hi⟩, by rw [← hi, ← mul_apply, ← zpow_one_add, add_comm, zpow_add_one, mul_apply, (f ^ i).injective.eq_iff] -lemma same_cycle_apply : same_cycle f x (f y) ↔ same_cycle f x y := -⟨λ ⟨i, hi⟩, ⟨-1 + i, by rw [zpow_add, mul_apply, hi, zpow_neg_one, inv_apply_self]⟩, - λ ⟨i, hi⟩, ⟨1 + i, by rw [zpow_add, mul_apply, hi, zpow_one]⟩⟩ +lemma same_cycle.eq_of_left (h : same_cycle f x y) (hx : is_fixed_pt f x) : x = y := +let ⟨n, hn⟩ := h in (hx.perm_zpow _).eq.symm.trans hn -lemma same_cycle_inv_apply : same_cycle f x (f⁻¹ y) ↔ same_cycle f x y := -by rw [← same_cycle_inv, same_cycle_apply, same_cycle_inv] +lemma same_cycle.eq_of_right (h : same_cycle f x y) (hy : is_fixed_pt f y) : x = y := +h.eq_of_left $ h.apply_eq_self_iff.2 hy -@[simp] lemma same_cycle_pow_left_iff {n : ℕ} : same_cycle f ((f ^ n) x) y ↔ same_cycle f x y := -begin - split, - { rintro ⟨k, rfl⟩, - use (k + n), - simp [zpow_add] }, - { rintro ⟨k, rfl⟩, - use (k - n), - rw [←zpow_coe_nat, ←mul_apply, ←zpow_add, int.sub_add_cancel] } -end +@[simp] lemma same_cycle_apply_left : same_cycle f (f x) y ↔ same_cycle f x y := +(equiv.add_right 1).exists_congr_left.trans $ by simp [zpow_sub, same_cycle] + +@[simp] lemma same_cycle_apply_right : same_cycle f x (f y) ↔ same_cycle f x y := +by rw [same_cycle_comm, same_cycle_apply_left, same_cycle_comm] + +@[simp] lemma same_cycle_inv_apply_left : same_cycle f (f⁻¹ x) y ↔ same_cycle f x y := +by rw [←same_cycle_apply_left, apply_inv_self] + +@[simp] lemma same_cycle_inv_apply_right : same_cycle f x (f⁻¹ y) ↔ same_cycle f x y := +by rw [←same_cycle_apply_right, apply_inv_self] @[simp] lemma same_cycle_zpow_left_iff {n : ℤ} : same_cycle f ((f ^ n) x) y ↔ same_cycle f x y := -begin - cases n, - { exact same_cycle_pow_left_iff }, - { rw [zpow_neg_succ_of_nat, ←inv_pow, ←same_cycle_inv, same_cycle_pow_left_iff, same_cycle_inv] } -end +(equiv.add_right (n : ℤ)).exists_congr_left.trans $ by simp [same_cycle, zpow_add] + +@[simp] lemma same_cycle_zpow_right_iff {n : ℤ} : same_cycle f x ((f ^ n) y) ↔ same_cycle f x y := +by rw [same_cycle_comm, same_cycle_zpow_left_iff, same_cycle_comm] + +@[simp] lemma same_cycle_pow_left_iff {n : ℕ} : same_cycle f ((f ^ n) x) y ↔ same_cycle f x y := +by rw [←zpow_coe_nat, same_cycle_zpow_left_iff] + +@[simp] lemma same_cycle_pow_right_iff {n : ℕ} : same_cycle f x ((f ^ n) y) ↔ same_cycle f x y := +by rw [←zpow_coe_nat, same_cycle_zpow_right_iff] + +alias same_cycle_pow_left_iff ↔ _ same_cycle.pow_left +alias same_cycle_pow_right_iff ↔ _ same_cycle.pow_right +alias same_cycle_zpow_left_iff ↔ _ same_cycle.zpow_left +alias same_cycle_zpow_right_iff ↔ _ same_cycle.zpow_right + +lemma same_cycle.of_pow {n : ℕ} : same_cycle (f ^ n) x y → same_cycle f x y := +λ ⟨m, h⟩, ⟨n * m, by simp [zpow_mul, h]⟩ + +lemma same_cycle.of_zpow {n : ℤ} : same_cycle (f ^ n) x y → same_cycle f x y := +λ ⟨m, h⟩, ⟨n * m, by simp [zpow_mul, h]⟩ + +@[simp] lemma same_cycle_subtype_perm {h} {x y : {x // p x}} : + (f.subtype_perm h).same_cycle x y ↔ f.same_cycle x y := +exists_congr $ λ n, by simp [subtype.ext_iff] + +alias same_cycle_subtype_perm ↔ _ same_cycle.subtype_perm + +@[simp] lemma same_cycle_extend_domain {p : β → Prop} [decidable_pred p] {f : α ≃ subtype p} : + same_cycle (g.extend_domain f) (f x) (f y) ↔ g.same_cycle x y := +exists_congr $ λ n, by rw [←extend_domain_zpow, extend_domain_apply_image, subtype.coe_inj, + f.injective.eq_iff] + +alias same_cycle_extend_domain ↔ _ same_cycle.extend_domain lemma same_cycle.nat' [finite α] : same_cycle f x y → ∃ i < order_of f, (f ^ i) x = y := begin @@ -132,71 +178,89 @@ end same_cycle -/ section is_cycle -variables [fintype α] [decidable_eq α] +variables {f g : perm α} {x y : α} -/-- A permutation is a cycle when any two nonfixed points of the permutation are related by repeated - application of the permutation. -/ -def is_cycle (f : perm β) : Prop := ∃ x, f x ≠ x ∧ ∀ y, f y ≠ y → ∃ i : ℤ, (f ^ i) x = y +/-- A cycle is a non identity permutation where any two nonfixed points of the permutation are +related by repeated application of the permutation. -/ +def is_cycle (f : perm α) : Prop := ∃ x, f x ≠ x ∧ ∀ ⦃y⦄, f y ≠ y → same_cycle f x y -lemma is_cycle.ne_one {f : perm β} (h : is_cycle f) : f ≠ 1 := -λ hf, by simpa [hf, is_cycle] using h +lemma is_cycle.ne_one (h : is_cycle f) : f ≠ 1 := λ hf, by simpa [hf, is_cycle] using h -@[simp] lemma not_is_cycle_one : ¬ (1 : perm β).is_cycle := -λ H, H.ne_one rfl +@[simp] lemma not_is_cycle_one : ¬ (1 : perm α).is_cycle := λ H, H.ne_one rfl -lemma is_cycle.two_le_card_support {f : perm α} (h : is_cycle f) : - 2 ≤ f.support.card := -two_le_card_support_of_ne_one h.ne_one +protected lemma is_cycle.same_cycle (hf : is_cycle f) (hx : f x ≠ x) (hy : f y ≠ y) : + same_cycle f x y := +let ⟨g, hg⟩ := hf in +let ⟨a, ha⟩ := hg.2 hx in +let ⟨b, hb⟩ := hg.2 hy in +⟨b - a, by rw [←ha, ←mul_apply, ←zpow_add, sub_add_cancel, hb]⟩ -lemma is_cycle_swap {α : Type*} [decidable_eq α] {x y : α} (hxy : x ≠ y) : is_cycle (swap x y) := -⟨y, by rwa swap_apply_right, - λ a (ha : ite (a = x) y (ite (a = y) x a) ≠ a), - if hya : y = a then ⟨0, hya⟩ - else ⟨1, by { rw [zpow_one, swap_apply_def], split_ifs at *; cc }⟩⟩ +lemma is_cycle.exists_zpow_eq : is_cycle f → f x ≠ x → f y ≠ y → ∃ i : ℤ, (f ^ i) x = y := +is_cycle.same_cycle + +lemma is_cycle.inv (hf : is_cycle f) : is_cycle f⁻¹ := +hf.imp $ λ x ⟨hx, h⟩, ⟨inv_eq_iff_eq.not.2 hx.symm, λ y hy, (h $ inv_eq_iff_eq.not.2 hy.symm).inv⟩ -lemma is_swap.is_cycle {α : Type*} [decidable_eq α] {f : perm α} (hf : is_swap f) : is_cycle f := +@[simp] lemma is_cycle_inv : is_cycle f⁻¹ ↔ is_cycle f := +⟨λ h, by { convert h.inv, rw inv_inv }, is_cycle.inv⟩ + +lemma is_cycle.conj : is_cycle f → is_cycle (g * f * g⁻¹) := begin - obtain ⟨x, y, hxy, rfl⟩ := hf, - exact is_cycle_swap hxy, + rintro ⟨x, hx, h⟩, + refine ⟨g x, by simp [coe_mul, inv_apply_self, hx], λ y hy, _⟩, + rw ←apply_inv_self g y, + exact (h $ eq_inv_iff_eq.not.2 hy).conj, end -lemma is_cycle.inv {f : perm β} (hf : is_cycle f) : is_cycle (f⁻¹) := -let ⟨x, hx⟩ := hf in -⟨x, by { simp only [inv_eq_iff_eq, *, forall_prop_of_true, ne.def] at *, cc }, - λ y hy, let ⟨i, hi⟩ := hx.2 y (by { simp only [inv_eq_iff_eq, *, forall_prop_of_true, - ne.def] at *, cc }) in - ⟨-i, by rwa [zpow_neg, inv_zpow, inv_inv]⟩⟩ - -lemma is_cycle.is_cycle_conj {f g : perm β} (hf : is_cycle f) : is_cycle (g * f * g⁻¹) := +protected lemma is_cycle.extend_domain {p : β → Prop} [decidable_pred p] (f : α ≃ subtype p) : + is_cycle g → is_cycle (g.extend_domain f) := begin - obtain ⟨a, ha1, ha2⟩ := hf, - refine ⟨g a, by simp [ha1], λ b hb, _⟩, - obtain ⟨i, hi⟩ := ha2 (g⁻¹ b) _, - { refine ⟨i, _⟩, - rw conj_zpow, - simp [hi] }, - { contrapose! hb, - rw [perm.mul_apply, perm.mul_apply, hb, apply_inv_self] } + rintro ⟨a, ha, ha'⟩, + refine ⟨f a, _, λ b hb, _⟩, + { rw extend_domain_apply_image, + exact subtype.coe_injective.ne (f.injective.ne ha) }, + have h : b = f (f.symm ⟨b, of_not_not $ hb ∘ extend_domain_apply_not_subtype _ _⟩), + { rw [apply_symm_apply, subtype.coe_mk] }, + rw h at ⊢ hb, + simp only [extend_domain_apply_image, subtype.coe_injective.ne_iff, f.injective.ne_iff] at hb, + exact (ha' hb).extend_domain, end -lemma is_cycle.exists_zpow_eq {f : perm β} (hf : is_cycle f) {x y : β} - (hx : f x ≠ x) (hy : f y ≠ y) : ∃ i : ℤ, (f ^ i) x = y := -let ⟨g, hg⟩ := hf in -let ⟨a, ha⟩ := hg.2 x hx in -let ⟨b, hb⟩ := hg.2 y hy in -⟨b - a, by rw [← ha, ← mul_apply, ← zpow_add, sub_add_cancel, hb]⟩ +lemma is_cycle_iff_same_cycle (hx : f x ≠ x) : is_cycle f ↔ ∀ {y}, same_cycle f x y ↔ f y ≠ y := +⟨λ hf y, ⟨λ ⟨i, hi⟩ hy, hx $ + by { rw [← zpow_apply_eq_self_of_apply_eq_self hy i, (f ^ i).injective.eq_iff] at hi, + rw [hi, hy] }, + hf.exists_zpow_eq hx⟩, + λ h, ⟨x, hx, λ y hy, h.2 hy⟩⟩ -lemma is_cycle.same_cycle {f : perm β} (hf : is_cycle f) {x y : β} (hx : f x ≠ x) (hy : f y ≠ y) : - same_cycle f x y := -hf.exists_zpow_eq hx hy +section finite +variables [finite α] -lemma is_cycle.exists_pow_eq [finite β] {f : perm β} (hf : is_cycle f) {x y : β} - (hx : f x ≠ x) (hy : f y ≠ y) : ∃ i : ℕ, (f ^ i) x = y := +lemma is_cycle.exists_pow_eq (hf : is_cycle f) (hx : f x ≠ x) (hy : f y ≠ y) : + ∃ i : ℕ, (f ^ i) x = y := let ⟨n, hn⟩ := hf.exists_zpow_eq hx hy in by classical; exact ⟨(n % order_of f).to_nat, by { have := n.mod_nonneg (int.coe_nat_ne_zero.mpr (ne_of_gt (order_of_pos f))), rwa [← zpow_coe_nat, int.to_nat_of_nonneg this, ← zpow_eq_mod_order_of] }⟩ +end finite + +variables [decidable_eq α] + +lemma is_cycle_swap (hxy : x ≠ y) : is_cycle (swap x y) := +⟨y, by rwa swap_apply_right, + λ a (ha : ite (a = x) y (ite (a = y) x a) ≠ a), + if hya : y = a then ⟨0, hya⟩ + else ⟨1, by { rw [zpow_one, swap_apply_def], split_ifs at *; cc }⟩⟩ + +protected lemma is_swap.is_cycle : is_swap f → is_cycle f := +by { rintro ⟨x, y, hxy, rfl⟩, exact is_cycle_swap hxy } + +variables [fintype α] + +lemma is_cycle.two_le_card_support (h : is_cycle f) : 2 ≤ f.support.card := +two_le_card_support_of_ne_one h.ne_one + lemma is_cycle.exists_pow_eq_one [finite β] {f : perm β} (hf : is_cycle f) : ∃ (k : ℕ) (hk : 1 < k), f ^ k = 1 := begin @@ -226,12 +290,12 @@ begin ext y, by_cases hy : σ y = y, { simp_rw [subtype.coe_mk, zpow_apply_eq_self_of_apply_eq_self hy] }, - { obtain ⟨i, rfl⟩ := (classical.some_spec hσ).2 y hy, + { obtain ⟨i, rfl⟩ := (classical.some_spec hσ).2 hy, rw [subtype.coe_mk, subtype.coe_mk, zpow_apply_comm σ m i, zpow_apply_comm σ n i], exact congr_arg _ (subtype.ext_iff.mp h) } }, by { rintros ⟨y, hy⟩, rw [finset.mem_coe, mem_support] at hy, - obtain ⟨n, rfl⟩ := (classical.some_spec hσ).2 y hy, + obtain ⟨n, rfl⟩ := (classical.some_spec hσ).2 hy, exact ⟨⟨σ ^ n, n, rfl⟩, rfl⟩ }, end @@ -246,10 +310,10 @@ rfl ⟨σ ^ n, n, rfl⟩ := (equiv.symm_apply_eq _).2 hσ.zpowers_equiv_support_apply -lemma order_of_is_cycle {σ : perm α} (hσ : is_cycle σ) : order_of σ = σ.support.card := +protected lemma is_cycle.order_of (hf : is_cycle f) : order_of f = f.support.card := begin rw [order_eq_card_zpowers, ←fintype.card_coe], - convert fintype.card_congr (is_cycle.zpowers_equiv_support hσ), + convert fintype.card_congr (is_cycle.zpowers_equiv_support hf), end lemma is_cycle_swap_mul_aux₁ {α : Type*} [decidable_eq α] : ∀ (n : ℕ) {b x : α} {f : perm α} @@ -299,13 +363,13 @@ lemma is_cycle.eq_swap_of_apply_apply_eq_self {α : Type*} [decidable_eq α] (hfx : f x ≠ x) (hffx : f (f x) = x) : f = swap x (f x) := equiv.ext $ λ y, let ⟨z, hz⟩ := hf in -let ⟨i, hi⟩ := hz.2 x hfx in +let ⟨i, hi⟩ := hz.2 hfx in if hyx : y = x then by simp [hyx] else if hfyx : y = f x then by simp [hfyx, hffx] else begin rw [swap_apply_of_ne_of_ne hyx hfyx], refine by_contradiction (λ hy, _), - cases hz.2 y hy with j hj, + cases hz.2 hy with j hj, rw [← sub_add_cancel j i, zpow_add, mul_apply, hi] at hj, cases zpow_apply_eq_of_apply_apply_eq_self hffx (j - i) with hji hji, { rw [← hj, hji] at hyx, cc }, @@ -349,70 +413,37 @@ calc sign f = sign (swap x (f x) * (swap x (f x) * f)) : pow_one, neg_mul_neg] } using_well_founded {rel_tac := λ _ _, `[exact ⟨_, measure_wf (λ f, f.support.card)⟩]} -lemma is_cycle_of_is_cycle_pow {σ : perm α} {n : ℕ} - (h1 : is_cycle (σ ^ n)) (h2 : σ.support ≤ (σ ^ n).support) : is_cycle σ := +lemma is_cycle.of_pow {n : ℕ} (h1 : is_cycle (f ^ n)) (h2 : f.support ⊆ (f ^ n).support) : + is_cycle f := begin - have key : ∀ x : α, (σ ^ n) x ≠ x ↔ σ x ≠ x, - { simp_rw [←mem_support], - exact finset.ext_iff.mp (le_antisymm (support_pow_le σ n) h2) }, + have key : ∀ x : α, (f ^ n) x ≠ x ↔ f x ≠ x, + { simp_rw [←mem_support, ←finset.ext_iff], + exact (support_pow_le _ n).antisymm h2 }, obtain ⟨x, hx1, hx2⟩ := h1, refine ⟨x, (key x).mp hx1, λ y hy, _⟩, - cases (hx2 y ((key y).mpr hy)) with i _, + cases (hx2 ((key y).mpr hy)) with i _, exact ⟨n * i, by rwa zpow_mul⟩ end -- The lemma `support_zpow_le` is relevant. It means that `h2` is equivalent to -- `σ.support = (σ ^ n).support`, as well as to `σ.support.card ≤ (σ ^ n).support.card`. -lemma is_cycle_of_is_cycle_zpow {σ : perm α} {n : ℤ} - (h1 : is_cycle (σ ^ n)) (h2 : σ.support ≤ (σ ^ n).support) : is_cycle σ := +lemma is_cycle.of_zpow {n : ℤ} (h1 : is_cycle (f ^ n)) (h2 : f.support ⊆ (f ^ n).support) : + is_cycle f := begin cases n, - { exact is_cycle_of_is_cycle_pow h1 h2 }, + { exact h1.of_pow h2 }, { simp only [le_eq_subset, zpow_neg_succ_of_nat, perm.support_inv] at h1 h2, - simpa using is_cycle_of_is_cycle_pow h1.inv h2 } + simpa using h1.inv.of_pow h2 } end -lemma is_cycle.extend_domain {α : Type*} {p : β → Prop} [decidable_pred p] - (f : α ≃ subtype p) {g : perm α} (h : is_cycle g) : - is_cycle (g.extend_domain f) := -begin - obtain ⟨a, ha, ha'⟩ := h, - refine ⟨f a, _, λ b hb, _⟩, - { rw extend_domain_apply_image, - exact λ con, ha (f.injective (subtype.coe_injective con)) }, - by_cases pb : p b, - { obtain ⟨i, hi⟩ := ha' (f.symm ⟨b, pb⟩) (λ con, hb _), - { refine ⟨i, _⟩, - have hnat : ∀ (k : ℕ) (a : α), (g.extend_domain f ^ k) ↑(f a) = f ((g ^ k) a), - { intros k a, - induction k with k ih, { refl }, - rw [pow_succ, perm.mul_apply, ih, extend_domain_apply_image, pow_succ, perm.mul_apply] }, - have hint : ∀ (k : ℤ) (a : α), (g.extend_domain f ^ k) ↑(f a) = f ((g ^ k) a), - { intros k a, - induction k with k k, - { rw [zpow_of_nat, zpow_of_nat, hnat] }, - rw [zpow_neg_succ_of_nat, zpow_neg_succ_of_nat, inv_eq_iff_eq, hnat, apply_inv_self] }, - rw [hint, hi, apply_symm_apply, subtype.coe_mk] }, - { rw [extend_domain_apply_subtype _ _ pb, con, apply_symm_apply, subtype.coe_mk] } }, - { exact (hb (extend_domain_apply_not_subtype _ _ pb)).elim } -end - -lemma same_cycle_cycle {f : perm β} {x : β} (hx : f x ≠ x) : is_cycle f ↔ - (∀ {y}, same_cycle f x y ↔ f y ≠ y) := -⟨λ hf y, ⟨λ ⟨i, hi⟩ hy, hx $ - by { rw [← zpow_apply_eq_self_of_apply_eq_self hy i, (f ^ i).injective.eq_iff] at hi, - rw [hi, hy] }, - hf.exists_zpow_eq hx⟩, - λ h, ⟨x, hx, λ y hy, h.2 hy⟩⟩ - lemma nodup_of_pairwise_disjoint_cycles {l : list (perm β)} (h1 : ∀ f ∈ l, is_cycle f) (h2 : l.pairwise disjoint) : l.nodup := nodup_of_pairwise_disjoint (λ h, (h1 1 h).ne_one rfl) h2 /-- Unlike `support_congr`, which assumes that `∀ (x ∈ g.support), f x = g x)`, here we have the weaker assumption that `∀ (x ∈ f.support), f x = g x`. -/ -lemma is_cycle.support_congr {f g : perm α} (hf : is_cycle f) (hg : is_cycle g) - (h : f.support ⊆ g.support) (h' : ∀ (x ∈ f.support), f x = g x) : f = g := +lemma is_cycle.support_congr (hf : is_cycle f) (hg : is_cycle g) (h : f.support ⊆ g.support) + (h' : ∀ x ∈ f.support, f x = g x) : f = g := begin have : f.support = g.support, { refine le_antisymm h _, @@ -432,9 +463,9 @@ end /-- If two cyclic permutations agree on all terms in their intersection, and that intersection is not empty, then the two cyclic permutations must be equal. -/ -lemma is_cycle.eq_on_support_inter_nonempty_congr {f g : perm α} - (hf : is_cycle f) (hg : is_cycle g) (h : ∀ (x ∈ f.support ∩ g.support), f x = g x) {x : α} - (hx : f x = g x) (hx' : x ∈ f.support) : f = g := +lemma is_cycle.eq_on_support_inter_nonempty_congr (hf : is_cycle f) (hg : is_cycle g) + (h : ∀ x ∈ f.support ∩ g.support, f x = g x) (hx : f x = g x) (hx' : x ∈ f.support) : + f = g := begin have hx'' : x ∈ g.support, { rwa [mem_support, ←hx, ←mem_support] }, @@ -446,7 +477,7 @@ begin exact hf.support_congr hg this h end -lemma is_cycle.support_pow_eq_iff {f : perm α} (hf : is_cycle f) {n : ℕ} : +lemma is_cycle.support_pow_eq_iff (hf : is_cycle f) {n : ℕ} : support (f ^ n) = support f ↔ ¬ order_of f ∣ n := begin rw order_of_dvd_iff_pow_eq_one, @@ -467,6 +498,11 @@ begin simpa using H } } end +lemma is_cycle.support_pow_of_pos_of_lt_order_of (hf : is_cycle f) {n : ℕ} (npos : 0 < n) + (hn : n < order_of f) : + (f ^ n).support = f.support := +hf.support_pow_eq_iff.2 $ nat.not_dvd_of_pos_of_lt npos hn + lemma is_cycle.pow_iff [finite β] {f : perm β} (hf : is_cycle f) {n : ℕ} : is_cycle (f ^ n) ↔ n.coprime (order_of f) := begin @@ -480,7 +516,7 @@ begin refine h.ne_one _, simp [pow_mul, pow_order_of_eq_one] }, have : order_of (f ^ n) = order_of f, - { rw [order_of_is_cycle h, hr, order_of_is_cycle hf] }, + { rw [h.order_of, hr, hf.order_of] }, rw [order_of_pow, nat.div_eq_self] at this, cases this, { exact absurd this (order_of_pos _).ne' }, @@ -488,12 +524,12 @@ begin { intro h, obtain ⟨m, hm⟩ := exists_pow_eq_self_of_coprime h, have hf' : is_cycle ((f ^ n) ^ m) := by rwa hm, - refine is_cycle_of_is_cycle_pow hf' _, - intros x hx, + refine hf'.of_pow (λ x hx, _), rw [hm], exact support_pow_le _ n hx } end +-- TODO: Define a `set`-valued support to get rid of the `finite β` assumption lemma is_cycle.pow_eq_one_iff [finite β] {f : perm β} (hf : is_cycle f) {n : ℕ} : f ^ n = 1 ↔ ∃ x, f x ≠ x ∧ (f ^ n) x = x := begin @@ -512,6 +548,19 @@ begin rw [pow_mul, pow_order_of_eq_one, one_pow] } } end +-- TODO: Define a `set`-valued support to get rid of the `finite β` assumption +lemma is_cycle.pow_eq_one_iff' [finite β] {f : perm β} (hf : is_cycle f) {n : ℕ} {x : β} + (hx : f x ≠ x) : + f ^ n = 1 ↔ (f ^ n) x = x := +⟨λ h, fun_like.congr_fun h x, λ h, hf.pow_eq_one_iff.2 ⟨x, hx, h⟩⟩ + +-- TODO: Define a `set`-valued support to get rid of the `finite β` assumption +lemma is_cycle.pow_eq_one_iff'' [finite β] {f : perm β} (hf : is_cycle f) {n : ℕ} : + f ^ n = 1 ↔ ∀ x, f x ≠ x → (f ^ n) x = x := +⟨λ h x hx, (hf.pow_eq_one_iff' hx).1 h, λ h, let ⟨x, hx, _⟩ := id hf in + (hf.pow_eq_one_iff' hx).2 (h _ hx)⟩ + +-- TODO: Define a `set`-valued support to get rid of the `finite β` assumption lemma is_cycle.pow_eq_pow_iff [finite β] {f : perm β} (hf : is_cycle f) {a b : ℕ} : f ^ a = f ^ b ↔ ∃ x, f x ≠ x ∧ (f ^ a) x = (f ^ b) x := begin @@ -537,15 +586,6 @@ begin contradiction }} end -lemma is_cycle.mem_support_pos_pow_iff_of_lt_order_of {f : perm α} (hf : is_cycle f) - {n : ℕ} (npos : 0 < n) (hn : n < order_of f) {x : α} : - x ∈ (f ^ n).support ↔ x ∈ f.support := -begin - have : ¬ order_of f ∣ n := nat.not_dvd_of_pos_of_lt npos hn, - rw ←hf.support_pow_eq_iff at this, - rw this -end - lemma is_cycle.is_cycle_pow_pos_of_lt_prime_order [finite β] {f : perm β} (hf : is_cycle f) (hf' : (order_of f).prime) (n : ℕ) (hn : 0 < n) (hn' : n < order_of f) : is_cycle (f ^ n) := begin @@ -558,25 +598,25 @@ begin obtain ⟨m, hm⟩ := exists_pow_eq_self_of_coprime this, have hf'' := hf, rw ←hm at hf'', - refine is_cycle_of_is_cycle_pow hf'' _, + refine hf''.of_pow _, rw [hm], exact support_pow_le f n end end is_cycle - /-! ### `cycle_of` -/ -variables [decidable_eq α] +section cycle_of +variables [decidable_eq α] [fintype α] {f g : perm α} {x y : α} /-- `f.cycle_of x` is the cycle of the permutation `f` to which `x` belongs. -/ -def cycle_of [fintype α] (f : perm α) (x : α) : perm α := -of_subtype (subtype_perm f (λ _, same_cycle_apply.symm) : perm {y // same_cycle f x y}) +def cycle_of (f : perm α) (x : α) : perm α := +of_subtype (subtype_perm f (λ _, same_cycle_apply_right.symm) : perm {y // same_cycle f x y}) -lemma cycle_of_apply [fintype α] (f : perm α) (x y : α) : +lemma cycle_of_apply (f : perm α) (x y : α) : cycle_of f x y = if same_cycle f x y then f y else y := begin dsimp only [cycle_of], @@ -585,40 +625,32 @@ begin { apply of_subtype_apply_of_not_mem, exact h }, end -lemma cycle_of_inv [fintype α] (f : perm α) (x : α) : - (cycle_of f x)⁻¹ = cycle_of f⁻¹ x := +lemma cycle_of_inv (f : perm α) (x : α) : (cycle_of f x)⁻¹ = cycle_of f⁻¹ x := equiv.ext $ λ y, begin rw [inv_eq_iff_eq, cycle_of_apply, cycle_of_apply], - split_ifs; simp [*, same_cycle_inv, same_cycle_inv_apply] at * + split_ifs; simp [*, same_cycle_inv, same_cycle_inv_apply_right] at * end -@[simp] lemma cycle_of_pow_apply_self [fintype α] (f : perm α) (x : α) : +@[simp] lemma cycle_of_pow_apply_self (f : perm α) (x : α) : ∀ n : ℕ, (cycle_of f x ^ n) x = (f ^ n) x | 0 := rfl | (n+1) := by { rw [pow_succ, mul_apply, cycle_of_apply, cycle_of_pow_apply_self, if_pos, pow_succ, mul_apply], exact ⟨n, rfl⟩ } -@[simp] lemma cycle_of_zpow_apply_self [fintype α] (f : perm α) (x : α) : +@[simp] lemma cycle_of_zpow_apply_self (f : perm α) (x : α) : ∀ n : ℤ, (cycle_of f x ^ n) x = (f ^ n) x | (n : ℕ) := cycle_of_pow_apply_self f x n | -[1+ n] := by rw [zpow_neg_succ_of_nat, ← inv_pow, cycle_of_inv, zpow_neg_succ_of_nat, ← inv_pow, cycle_of_pow_apply_self] -lemma same_cycle.cycle_of_apply [fintype α] {f : perm α} {x y : α} (h : same_cycle f x y) : - cycle_of f x y = f y := -begin - apply of_subtype_apply_of_mem, exact h, -end +lemma same_cycle.cycle_of_apply : same_cycle f x y → cycle_of f x y = f y := +of_subtype_apply_of_mem _ -lemma cycle_of_apply_of_not_same_cycle [fintype α] {f : perm α} {x y : α} (h : ¬same_cycle f x y) : - cycle_of f x y = y := -begin - apply of_subtype_apply_of_not_mem, exact h, -end +lemma cycle_of_apply_of_not_same_cycle : ¬ same_cycle f x y → cycle_of f x y = y := +of_subtype_apply_of_not_mem _ -lemma same_cycle.cycle_of_eq [fintype α] {f : perm α} {x y : α} (h : same_cycle f x y) : - cycle_of f x = cycle_of f y := +lemma same_cycle.cycle_of_eq (h : same_cycle f x y) : cycle_of f x = cycle_of f y := begin ext z, rw cycle_of_apply, @@ -627,7 +659,7 @@ begin { exact (cycle_of_apply_of_not_same_cycle (mt h.trans hz)).symm } end -@[simp] lemma cycle_of_apply_apply_zpow_self [fintype α] (f : perm α) (x : α) (k : ℤ) : +@[simp] lemma cycle_of_apply_apply_zpow_self (f : perm α) (x : α) (k : ℤ) : cycle_of f x ((f ^ k) x) = (f ^ (k + 1)) x := begin rw same_cycle.cycle_of_apply, @@ -635,24 +667,22 @@ begin { exact ⟨k, rfl⟩ } end -@[simp] lemma cycle_of_apply_apply_pow_self [fintype α] (f : perm α) (x : α) (k : ℕ) : +@[simp] lemma cycle_of_apply_apply_pow_self (f : perm α) (x : α) (k : ℕ) : cycle_of f x ((f ^ k) x) = (f ^ (k + 1)) x := by convert cycle_of_apply_apply_zpow_self f x k using 1 -@[simp] lemma cycle_of_apply_apply_self [fintype α] (f : perm α) (x : α) : - cycle_of f x (f x) = f (f x) := +@[simp] lemma cycle_of_apply_apply_self (f : perm α) (x : α) : cycle_of f x (f x) = f (f x) := by convert cycle_of_apply_apply_pow_self f x 1 using 1 -@[simp] lemma cycle_of_apply_self [fintype α] (f : perm α) (x : α) : - cycle_of f x x = f x := (same_cycle.refl _ _).cycle_of_apply +@[simp] lemma cycle_of_apply_self (f : perm α) (x : α) : cycle_of f x x = f x := +same_cycle.rfl.cycle_of_apply -lemma is_cycle.cycle_of_eq [fintype α] {f : perm α} (hf : is_cycle f) {x : α} (hx : f x ≠ x) : - cycle_of f x = f := +lemma is_cycle.cycle_of_eq (hf : is_cycle f) (hx : f x ≠ x) : cycle_of f x = f := equiv.ext $ λ y, - if h : same_cycle f x y then by rw [h.cycle_of_apply] - else by rw [cycle_of_apply_of_not_same_cycle h, not_not.1 (mt ((same_cycle_cycle hx).1 hf).2 h)] + if h : same_cycle f x y then by rw [h.cycle_of_apply] else + by rw [cycle_of_apply_of_not_same_cycle h, not_not.1 (mt ((is_cycle_iff_same_cycle hx).1 hf).2 h)] -@[simp] lemma cycle_of_eq_one_iff [fintype α] (f : perm α) {x : α} : cycle_of f x = 1 ↔ f x = x := +@[simp] lemma cycle_of_eq_one_iff (f : perm α) : cycle_of f x = 1 ↔ f x = x := begin simp_rw [ext_iff, cycle_of_apply, one_apply], refine ⟨λ h, (if_pos (same_cycle.refl f x)).symm.trans (h x), λ h y, _⟩, @@ -661,40 +691,36 @@ begin { exact if_neg (mt same_cycle.apply_eq_self_iff (by tauto)) }, end -@[simp] lemma cycle_of_self_apply [fintype α] (f : perm α) (x : α) : - cycle_of f (f x) = cycle_of f x := -(same_cycle_apply.mpr (same_cycle.refl _ _)).symm.cycle_of_eq +@[simp] lemma cycle_of_self_apply (f : perm α) (x : α) : cycle_of f (f x) = cycle_of f x := +(same_cycle_apply_right.2 same_cycle.rfl).symm.cycle_of_eq -@[simp] lemma cycle_of_self_apply_pow [fintype α] (f : perm α) (n : ℕ) (x : α) : +@[simp] lemma cycle_of_self_apply_pow (f : perm α) (n : ℕ) (x : α) : cycle_of f ((f ^ n) x) = cycle_of f x := -(same_cycle_pow_left_iff.mpr (same_cycle.refl _ _)).cycle_of_eq +same_cycle.rfl.pow_left.cycle_of_eq -@[simp] lemma cycle_of_self_apply_zpow [fintype α] (f : perm α) (n : ℤ) (x : α) : +@[simp] lemma cycle_of_self_apply_zpow (f : perm α) (n : ℤ) (x : α) : cycle_of f ((f ^ n) x) = cycle_of f x := -(same_cycle_zpow_left_iff.mpr (same_cycle.refl _ _)).cycle_of_eq +same_cycle.rfl.zpow_left.cycle_of_eq -lemma is_cycle.cycle_of [fintype α] {f : perm α} (hf : is_cycle f) {x : α} : - cycle_of f x = if f x = x then 1 else f := +protected lemma is_cycle.cycle_of (hf : is_cycle f) : cycle_of f x = if f x = x then 1 else f := begin by_cases hx : f x = x, { rwa [if_pos hx, cycle_of_eq_one_iff] }, { rwa [if_neg hx, hf.cycle_of_eq] }, end -lemma cycle_of_one [fintype α] (x : α) : cycle_of 1 x = 1 := -(cycle_of_eq_one_iff 1).mpr rfl +lemma cycle_of_one (x : α) : cycle_of 1 x = 1 := (cycle_of_eq_one_iff 1).mpr rfl -lemma is_cycle_cycle_of [fintype α] (f : perm α) {x : α} (hx : f x ≠ x) : is_cycle (cycle_of f x) := -have cycle_of f x x ≠ x, by rwa [(same_cycle.refl _ _).cycle_of_apply], -(same_cycle_cycle this).2 $ λ y, +lemma is_cycle_cycle_of (f : perm α) (hx : f x ≠ x) : is_cycle (cycle_of f x) := +have cycle_of f x x ≠ x, by rwa [same_cycle.rfl.cycle_of_apply], +(is_cycle_iff_same_cycle this).2 $ λ y, ⟨λ h, mt h.apply_eq_self_iff.2 this, λ h, if hxy : same_cycle f x y then let ⟨i, hi⟩ := hxy in ⟨i, by rw [cycle_of_zpow_apply_self, hi]⟩ else by { rw [cycle_of_apply_of_not_same_cycle hxy] at h, exact (h rfl).elim }⟩ -@[simp] lemma two_le_card_support_cycle_of_iff [fintype α] {f : perm α} {x : α} : - 2 ≤ card (cycle_of f x).support ↔ f x ≠ x := +@[simp] lemma two_le_card_support_cycle_of_iff : 2 ≤ card (cycle_of f x).support ↔ f x ≠ x := begin refine ⟨λ h, _, λ h, by simpa using (is_cycle_cycle_of _ h).two_le_card_support⟩, contrapose! h, @@ -702,19 +728,18 @@ begin simp [h] end -@[simp] lemma card_support_cycle_of_pos_iff [fintype α] {f : perm α} {x : α} : - 0 < card (cycle_of f x).support ↔ f x ≠ x := +@[simp] lemma card_support_cycle_of_pos_iff : 0 < card (cycle_of f x).support ↔ f x ≠ x := begin rw [←two_le_card_support_cycle_of_iff, ←nat.succ_le_iff], exact ⟨λ h, or.resolve_left h.eq_or_lt (card_support_ne_one _).symm, zero_lt_two.trans_le⟩ end -lemma pow_apply_eq_pow_mod_order_of_cycle_of_apply [fintype α] (f : perm α) (n : ℕ) (x : α) : +lemma pow_apply_eq_pow_mod_order_of_cycle_of_apply (f : perm α) (n : ℕ) (x : α) : (f ^ n) x = (f ^ (n % order_of (cycle_of f x))) x := by rw [←cycle_of_pow_apply_self f, ←cycle_of_pow_apply_self f, pow_eq_mod_order_of] -lemma cycle_of_mul_of_apply_right_eq_self [fintype α] {f g : perm α} - (h : _root_.commute f g) (x : α) (hx : g x = x) : (f * g).cycle_of x = f.cycle_of x := +lemma cycle_of_mul_of_apply_right_eq_self (h : _root_.commute f g) (x : α) (hx : g x = x) : + (f * g).cycle_of x = f.cycle_of x := begin ext y, by_cases hxy : (f * g).same_cycle x y, @@ -728,7 +753,7 @@ begin simp [h.mul_zpow, zpow_apply_eq_self_of_apply_eq_self hx] } end -lemma disjoint.cycle_of_mul_distrib [fintype α] {f g : perm α} (h : f.disjoint g) (x : α) : +lemma disjoint.cycle_of_mul_distrib (h : f.disjoint g) (x : α) : (f * g).cycle_of x = (f.cycle_of x * g.cycle_of x) := begin cases (disjoint_iff_eq_or_eq.mp h) x with hfx hgx, @@ -736,12 +761,9 @@ begin { simp [cycle_of_mul_of_apply_right_eq_self h.commute, hgx] } end -lemma support_cycle_of_eq_nil_iff [fintype α] {f : perm α} {x : α} : - (f.cycle_of x).support = ∅ ↔ x ∉ f.support := -by simp +lemma support_cycle_of_eq_nil_iff : (f.cycle_of x).support = ∅ ↔ x ∉ f.support := by simp -lemma support_cycle_of_le [fintype α] (f : perm α) (x : α) : - support (f.cycle_of x) ≤ support f := +lemma support_cycle_of_le (f : perm α) (x : α) : support (f.cycle_of x) ≤ support f := begin intros y hy, rw [mem_support, cycle_of_apply] at hy, @@ -750,8 +772,7 @@ begin { exact absurd rfl hy } end -lemma mem_support_cycle_of_iff [fintype α] {f : perm α} {x y : α} : - y ∈ support (f.cycle_of x) ↔ same_cycle f x y ∧ x ∈ support f := +lemma mem_support_cycle_of_iff : y ∈ support (f.cycle_of x) ↔ same_cycle f x y ∧ x ∈ support f := begin by_cases hx : f x = x, { rw (cycle_of_eq_one_iff _).mpr hx, @@ -765,23 +786,24 @@ begin { simpa [hx] using hy } } end -lemma same_cycle.mem_support_iff [fintype α] {f : perm α} {x y : α} (h : same_cycle f x y) : - x ∈ support f ↔ y ∈ support f := +lemma mem_support_cycle_of_iff' (hx : f x ≠ x) : y ∈ support (f.cycle_of x) ↔ same_cycle f x y := +by rw [mem_support_cycle_of_iff, and_iff_left (mem_support.2 hx)] + +lemma same_cycle.mem_support_iff (h : same_cycle f x y) : x ∈ support f ↔ y ∈ support f := ⟨λ hx, support_cycle_of_le f x (mem_support_cycle_of_iff.mpr ⟨h, hx⟩), λ hy, support_cycle_of_le f y (mem_support_cycle_of_iff.mpr ⟨h.symm, hy⟩)⟩ -lemma pow_mod_card_support_cycle_of_self_apply [fintype α] (f : perm α) (n : ℕ) (x : α) : +lemma pow_mod_card_support_cycle_of_self_apply (f : perm α) (n : ℕ) (x : α) : (f ^ (n % (f.cycle_of x).support.card)) x = (f ^ n) x := begin by_cases hx : f x = x, { rw [pow_apply_eq_self_of_apply_eq_self hx, pow_apply_eq_self_of_apply_eq_self hx] }, - { rw [←cycle_of_pow_apply_self, ←cycle_of_pow_apply_self f, - ←order_of_is_cycle (is_cycle_cycle_of f hx), ←pow_eq_mod_order_of] } + { rw [←cycle_of_pow_apply_self, ←cycle_of_pow_apply_self f, ←(is_cycle_cycle_of f hx).order_of, + ←pow_eq_mod_order_of] } end /-- x is in the support of f iff cycle_of f x is a cycle.-/ -lemma is_cycle_cycle_of_iff [fintype α] (f : perm α) {x : α} : - is_cycle (cycle_of f x) ↔ (f x ≠ x) := +lemma is_cycle_cycle_of_iff (f : perm α) : is_cycle (cycle_of f x) ↔ (f x ≠ x) := begin split, { intro hx, rw ne.def, rw ← cycle_of_eq_one_iff f, @@ -790,11 +812,14 @@ begin apply equiv.perm.is_cycle_cycle_of, exact hx } end +end cycle_of /-! ### `cycle_factors` -/ +variables [decidable_eq α] + /-- Given a list `l : list α` and a permutation `f : perm α` whose nonfixed points are all in `l`, recursively factors `f` into cycles. -/ def cycle_factors_aux [fintype α] : Π (l : list α) (f : perm α), @@ -823,7 +848,7 @@ else let ⟨m, hm₁, hm₂, hm₃⟩ := cycle_factors_aux l ((cycle_of f x)⁻ (list.pairwise_cons.1 ((hgm.pairwise_iff (λ a b (h : disjoint a b), h.symm)).2 hm₃)).1, classical.by_cases id $ λ hgy : g y ≠ y, (disjoint_prod_right _ this y).resolve_right $ - have hsc : same_cycle f⁻¹ x (f y), by rwa [same_cycle_inv, same_cycle_apply], + have hsc : same_cycle f⁻¹ x (f y), by rwa [same_cycle_inv, same_cycle_apply_right], by { rw [disjoint_prod_perm hm₃ hgm.symm, list.prod_cons, ← eq_inv_mul_iff_mul_eq] at hm₁, rwa [hm₁, mul_apply, mul_apply, cycle_of_inv, hsc.cycle_of_apply, @@ -1156,7 +1181,7 @@ begin { simp }, { intros g hg H hx, rw mem_support at hx, - rw [hg.cycle_of_eq hx, ←order_of_is_cycle hg], + rw [hg.cycle_of_eq hx, ←hg.order_of], exact H.nat' }, { rintros g h hd hg IH IH' ⟨m, rfl⟩ hx, cases (disjoint_iff_eq_or_eq.mp hd) x with hgx hhx, @@ -1202,7 +1227,7 @@ begin simpa using hx }, { simp only [perm.coe_one, id.def, pow_zero] at hk', subst hk', - rw [←order_of_is_cycle (is_cycle_cycle_of _ (mem_support.mp hx)), + rw [←(is_cycle_cycle_of _ $ mem_support.1 hx).order_of, ←cycle_of_pow_apply_self, pow_order_of_eq_one, one_apply] } }, { exact ⟨k + 1, by simp, nat.le_succ_of_le hk.le, hk'⟩ } }, { refine ⟨1, zero_lt_one, by simp, _⟩, @@ -1238,7 +1263,7 @@ begin induction n with n ih, { exact subset_closure (set.mem_insert_of_mem _ (set.mem_singleton _)) }, { convert H.mul_mem (H.mul_mem h3 ih) (H.inv_mem h3), - rw [mul_swap_eq_swap_mul, mul_inv_cancel_right], refl } }, + simp_rw [mul_swap_eq_swap_mul, mul_inv_cancel_right, pow_succ], refl } }, have step2 : ∀ (n : ℕ), swap x ((σ ^ n) x) ∈ H, { intro n, induction n with n ih, @@ -1277,11 +1302,11 @@ lemma closure_cycle_coprime_swap {n : ℕ} {σ : perm α} (h0 : nat.coprime n (f (h1 : is_cycle σ) (h2 : σ.support = finset.univ) (x : α) : closure ({σ, swap x ((σ ^ n) x)} : set (perm α)) = ⊤ := begin - rw [←finset.card_univ, ←h2, ←order_of_is_cycle h1] at h0, + rw [←finset.card_univ, ←h2, ←h1.order_of] at h0, cases exists_pow_eq_self_of_coprime h0 with m hm, have h2' : (σ ^ n).support = ⊤ := eq.trans (support_pow_coprime h0) h2, have h1' : is_cycle ((σ ^ n) ^ (m : ℤ)) := by rwa ← hm at h1, - replace h1' : is_cycle (σ ^ n) := is_cycle_of_is_cycle_pow h1' + replace h1' : is_cycle (σ ^ n) := h1'.of_pow (le_trans (support_pow_le σ n) (ge_of_eq (congr_arg support hm))), rw [eq_top_iff, ←closure_cycle_adjacent_swap h1' h2' x, closure_le, set.insert_subset], exact ⟨subgroup.pow_mem (closure _) (subset_closure (set.mem_insert σ _)) n, @@ -1299,7 +1324,7 @@ begin refine closure_cycle_coprime_swap (nat.coprime.symm (h0.coprime_iff_not_dvd.mpr (λ h, h4 _))) h1 h2 x, cases h with m hm, - rwa [hm, pow_mul, ←finset.card_univ, ←h2, ←order_of_is_cycle h1, + rwa [hm, pow_mul, ←finset.card_univ, ←h2, ←h1.order_of, pow_order_of_eq_one, one_pow, one_apply] at hi, end @@ -1329,10 +1354,8 @@ end theorem is_cycle.is_conj (hσ : is_cycle σ) (hτ : is_cycle τ) (h : σ.support.card = τ.support.card) : is_conj σ τ := begin - refine is_conj_of_support_equiv (hσ.zpowers_equiv_support.symm.trans - ((zpowers_equiv_zpowers begin - rw [order_of_is_cycle hσ, h, order_of_is_cycle hτ], - end).trans hτ.zpowers_equiv_support)) _, + refine is_conj_of_support_equiv (hσ.zpowers_equiv_support.symm.trans $ + (zpowers_equiv_zpowers $ by rw [hσ.order_of, h, hτ.order_of]).trans hτ.zpowers_equiv_support) _, intros x hx, simp only [perm.mul_apply, equiv.trans_apply, equiv.sum_congr_apply], obtain ⟨n, rfl⟩ := hσ.exists_pow_eq (classical.some_spec hσ).1 (mem_support.1 hx), diff --git a/src/group_theory/perm/cycle/concrete.lean b/src/group_theory/perm/cycle/concrete.lean index beceaf7abf0cc..1bc2a9449fa59 100644 --- a/src/group_theory/perm/cycle/concrete.lean +++ b/src/group_theory/perm/cycle/concrete.lean @@ -274,16 +274,16 @@ begin have hc : is_cycle (cycle_of p x) := is_cycle_cycle_of p hx, rw nodup_iff_nth_le_inj, rintros n m hn hm, - rw [length_to_list, ←order_of_is_cycle hc] at hm hn, + rw [length_to_list, ←hc.order_of] at hm hn, rw [←cycle_of_apply_self, ←ne.def, ←mem_support] at hx, rw [nth_le_to_list, nth_le_to_list, ←cycle_of_pow_apply_self p x n, ←cycle_of_pow_apply_self p x m], cases n; cases m, { simp }, - { rw [←hc.mem_support_pos_pow_iff_of_lt_order_of m.zero_lt_succ hm, + { rw [←hc.support_pow_of_pos_of_lt_order_of m.zero_lt_succ hm, mem_support, cycle_of_pow_apply_self] at hx, simp [hx.symm] }, - { rw [←hc.mem_support_pos_pow_iff_of_lt_order_of n.zero_lt_succ hn, + { rw [←hc.support_pow_of_pos_of_lt_order_of n.zero_lt_succ hn, mem_support, cycle_of_pow_apply_self] at hx, simp [hx] }, intro h, @@ -309,7 +309,7 @@ begin rw ←nth_le_to_list p x k (by simpa using hk) at hk', simp_rw ←hk', rw [next_nth_le _ (nodup_to_list _ _), nth_le_to_list, nth_le_to_list, ←mul_apply, ←pow_succ, - length_to_list, pow_apply_eq_pow_mod_order_of_cycle_of_apply p (k + 1), order_of_is_cycle], + length_to_list, pow_apply_eq_pow_mod_order_of_cycle_of_apply p (k + 1), is_cycle.order_of], exact is_cycle_cycle_of _ (mem_support.mp hy.right) end diff --git a/src/group_theory/perm/cycle/type.lean b/src/group_theory/perm/cycle/type.lean index 52bb7c485ca3d..3b66c3bc4cd65 100644 --- a/src/group_theory/perm/cycle/type.lean +++ b/src/group_theory/perm/cycle/type.lean @@ -130,7 +130,7 @@ begin { intro, simp }, { intros σ hσ τ, - rw [hσ.cycle_type, hσ.is_cycle_conj.cycle_type, card_support_conj] }, + rw [hσ.cycle_type, hσ.conj.cycle_type, card_support_conj] }, { intros σ τ hd hc hσ hτ π, rw [← conj_mul, hd.cycle_type, disjoint.cycle_type, hσ, hτ], intro a, @@ -178,7 +178,7 @@ cycle_induction_on lemma lcm_cycle_type (σ : perm α) : σ.cycle_type.lcm = order_of σ := cycle_induction_on (λ τ : perm α, τ.cycle_type.lcm = order_of τ) σ (by rw [cycle_type_one, lcm_zero, order_of_one]) - (λ σ hσ, by rw [hσ.cycle_type, coe_singleton, lcm_singleton, order_of_is_cycle hσ, + (λ σ hσ, by rw [hσ.cycle_type, coe_singleton, lcm_singleton, hσ.order_of, normalize_eq]) (λ σ τ hστ hc hσ hτ, by rw [hστ.cycle_type, lcm_add, lcm_eq_nat_lcm, hστ.order_of, hσ, hτ]) @@ -198,7 +198,7 @@ begin rw [cycle_type, multiset.mem_map], refine ⟨f.cycle_of x, _, _⟩, { rwa [←finset.mem_def, cycle_of_mem_cycle_factors_finset_iff, mem_support] }, - { simp [order_of_is_cycle (is_cycle_cycle_of _ hx)] } } + { simp [(is_cycle_cycle_of _ hx).order_of] } } end lemma two_dvd_card_support {σ : perm α} (hσ : σ ^ 2 = 1) : 2 ∣ σ.support.card := @@ -512,7 +512,7 @@ begin have hσ1 : order_of (σ : perm α) = fintype.card α := (order_of_subgroup σ).trans hσ, have hσ2 : is_cycle ↑σ := is_cycle_of_prime_order'' h0 hσ1, have hσ3 : (σ : perm α).support = ⊤ := - finset.eq_univ_of_card (σ : perm α).support ((order_of_is_cycle hσ2).symm.trans hσ1), + finset.eq_univ_of_card (σ : perm α).support (hσ2.order_of.symm.trans hσ1), have hσ4 : subgroup.closure {↑σ, τ} = ⊤ := closure_prime_cycle_swap h0 hσ2 hσ3 h3, rw [eq_top_iff, ←hσ4, subgroup.closure_le, set.insert_subset, set.singleton_subset_iff], exact ⟨subtype.mem σ, h2⟩, From 5a3e819569b0f12cbec59d740a2613018e7b8eec Mon Sep 17 00:00:00 2001 From: sgouezel Date: Thu, 5 Jan 2023 07:04:00 +0000 Subject: [PATCH 0142/1291] feat(order/filter/basic): put pp_nodot attribute on filter.tendsto (#18062) Make sure that `tendsto f p q` is not pretty-printed as `p.tendsto f q`, since this is not a property of `p` but rather of `f`. --- src/order/filter/basic.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/order/filter/basic.lean b/src/order/filter/basic.lean index 79538fcb2be36..e4e33dbfccce4 100644 --- a/src/order/filter/basic.lean +++ b/src/order/filter/basic.lean @@ -2366,7 +2366,7 @@ end list_traverse /-- `tendsto` is the generic "limit of a function" predicate. `tendsto f l₁ l₂` asserts that for every `l₂` neighborhood `a`, the `f`-preimage of `a` is an `l₁` neighborhood. -/ -def tendsto (f : α → β) (l₁ : filter α) (l₂ : filter β) := l₁.map f ≤ l₂ +@[pp_nodot] def tendsto (f : α → β) (l₁ : filter α) (l₂ : filter β) := l₁.map f ≤ l₂ lemma tendsto_def {f : α → β} {l₁ : filter α} {l₂ : filter β} : tendsto f l₁ l₂ ↔ ∀ s ∈ l₂, f ⁻¹' s ∈ l₁ := iff.rfl From 43a289b563d533cfb00f58b812d222371f6f6fac Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 5 Jan 2023 10:03:21 +0000 Subject: [PATCH 0143/1291] fix(data/finset/basic): correct the name of `finset.range_coe` which does not mention coercions (#18052) This also adds the missing lemma that should have had that name. --- src/algebra/big_operators/basic.lean | 2 +- src/data/finset/basic.lean | 5 ++++- src/data/multiset/fintype.lean | 2 +- src/data/nat/periodic.lean | 2 +- src/ring_theory/polynomial/vieta.lean | 2 +- src/ring_theory/witt_vector/mul_coeff.lean | 4 ++-- 6 files changed, 10 insertions(+), 7 deletions(-) diff --git a/src/algebra/big_operators/basic.lean b/src/algebra/big_operators/basic.lean index 8a7c4bbb5cd97..c6d418cdcec9b 100644 --- a/src/algebra/big_operators/basic.lean +++ b/src/algebra/big_operators/basic.lean @@ -1726,7 +1726,7 @@ lemma sup_powerset_len {α : Type*} [decidable_eq α] (x : multiset α) : finset.sup (finset.range (x.card + 1)) (λ k, x.powerset_len k) = x.powerset := begin convert bind_powerset_len x, - rw [multiset.bind, multiset.join, ←finset.range_coe, ←finset.sum_eq_multiset_sum], + rw [multiset.bind, multiset.join, ←finset.range_val, ←finset.sum_eq_multiset_sum], exact eq.symm (finset_sum_eq_sup_iff_disjoint.mpr (λ _ _ _ _ h, pairwise_disjoint_powerset_len x h)), end diff --git a/src/data/finset/basic.lean b/src/data/finset/basic.lean index 12e040fb99cf4..771c2eeb7e476 100644 --- a/src/data/finset/basic.lean +++ b/src/data/finset/basic.lean @@ -1925,10 +1925,13 @@ variables {n m l : ℕ} /-- `range n` is the set of natural numbers less than `n`. -/ def range (n : ℕ) : finset ℕ := ⟨_, nodup_range n⟩ -@[simp] theorem range_coe (n : ℕ) : (range n).1 = multiset.range n := rfl +@[simp] theorem range_val (n : ℕ) : (range n).1 = multiset.range n := rfl @[simp] theorem mem_range : m ∈ range n ↔ m < n := mem_range +@[simp, norm_cast] lemma coe_range (n : ℕ) : (range n : set ℕ) = set.Iio n := +set.ext $ λ _, mem_range + @[simp] theorem range_zero : range 0 = ∅ := rfl @[simp] theorem range_one : range 1 = {0} := rfl diff --git a/src/data/multiset/fintype.lean b/src/data/multiset/fintype.lean index 7f7e85dbc79dd..83b15e8cb94ad 100644 --- a/src/data/multiset/fintype.lean +++ b/src/data/multiset/fintype.lean @@ -175,7 +175,7 @@ end begin ext x, simp only [multiset.count_map, ← finset.filter_val, multiset.to_enum_finset_filter_eq, - finset.map_val, finset.range_coe, multiset.card_map, multiset.card_range], + finset.map_val, finset.range_val, multiset.card_map, multiset.card_range], end @[simp] lemma multiset.image_to_enum_finset_fst (m : multiset α) : diff --git a/src/data/nat/periodic.lean b/src/data/nat/periodic.lean index a5a14e738006c..501e32e15a159 100644 --- a/src/data/nat/periodic.lean +++ b/src/data/nat/periodic.lean @@ -40,7 +40,7 @@ lemma filter_multiset_Ico_card_eq_of_periodic (n a : ℕ) (p : ℕ → Prop) [de (pp : periodic p a) : (filter p (Ico n (n+a))).card = a.count p := begin - rw [count_eq_card_filter_range, finset.card, finset.filter_val, finset.range_coe, + rw [count_eq_card_filter_range, finset.card, finset.filter_val, finset.range_val, ←multiset_Ico_map_mod n, ←map_count_true_eq_filter_card, ←map_count_true_eq_filter_card, map_map, function.comp], simp only [pp.map_mod_nat], diff --git a/src/ring_theory/polynomial/vieta.lean b/src/ring_theory/polynomial/vieta.lean index e35b1f8358651..529f82e0f12cc 100644 --- a/src/ring_theory/polynomial/vieta.lean +++ b/src/ring_theory/polynomial/vieta.lean @@ -41,7 +41,7 @@ lemma prod_X_add_C_eq_sum_esymm (s : multiset R) : begin classical, rw [prod_map_add, antidiagonal_eq_map_powerset, map_map, ←bind_powerset_len, function.comp, - map_bind, sum_bind, finset.sum_eq_multiset_sum, finset.range_coe, map_congr (eq.refl _)], + map_bind, sum_bind, finset.sum_eq_multiset_sum, finset.range_val, map_congr (eq.refl _)], intros _ _, rw [esymm, ←sum_hom', ←sum_map_mul_right, map_congr (eq.refl _)], intros _ ht, diff --git a/src/ring_theory/witt_vector/mul_coeff.lean b/src/ring_theory/witt_vector/mul_coeff.lean index f59f173bd2bd6..196ef3ca87462 100644 --- a/src/ring_theory/witt_vector/mul_coeff.lean +++ b/src/ring_theory/witt_vector/mul_coeff.lean @@ -283,7 +283,7 @@ begin apply f₀, rintros ⟨a, ha⟩, apply function.uncurry (![x, y]), - simp only [true_and, multiset.mem_cons, range_coe, product_val, multiset.mem_range, + simp only [true_and, multiset.mem_cons, range_val, product_val, multiset.mem_range, multiset.mem_product, multiset.range_succ, mem_univ_val] at ha, refine ⟨a.fst, ⟨a.snd, _⟩⟩, cases ha with ha ha; linarith only [ha] }, @@ -296,7 +296,7 @@ begin ext a, cases a with a ha, cases a with i m, - simp only [true_and, multiset.mem_cons, range_coe, product_val, multiset.mem_range, + simp only [true_and, multiset.mem_cons, range_val, product_val, multiset.mem_range, multiset.mem_product, multiset.range_succ, mem_univ_val] at ha, have ha' : m < n + 1 := by cases ha with ha ha; linarith only [ha], fin_cases i; -- surely this case split is not necessary From 3c71503b4ee96c510917559a0cf06466b7cfe605 Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Thu, 5 Jan 2023 13:17:54 +0000 Subject: [PATCH 0144/1291] feat(data/finsupp/defs): add add_monoid_hom_class versions of map_range (#17738) - [x] depends on: #17717 Co-authored-by: Eric Wieser --- src/data/finsupp/defs.lean | 15 +++++++++++++++ 1 file changed, 15 insertions(+) diff --git a/src/data/finsupp/defs.lean b/src/data/finsupp/defs.lean index b0174f59c2437..ac9ae1845a73b 100644 --- a/src/data/finsupp/defs.lean +++ b/src/data/finsupp/defs.lean @@ -896,6 +896,11 @@ lemma map_range_add [add_zero_class N] map_range f hf (v₁ + v₂) = map_range f hf v₁ + map_range f hf v₂ := ext $ λ _, by simp only [hf', add_apply, map_range_apply] +lemma map_range_add' [add_zero_class N] [add_monoid_hom_class β M N] + {f : β} (v₁ v₂ : α →₀ M) : + map_range f (map_zero f) (v₁ + v₂) = map_range f (map_zero f) v₁ + map_range f (map_zero f) v₂ := +map_range_add (map_add f) v₁ v₂ + /-- Bundle `emb_domain f` as an additive map from `α →₀ M` to `β →₀ M`. -/ @[simps] def emb_domain.add_monoid_hom (f : α ↪ β) : (α →₀ M) →+ β →₀ M := { to_fun := λ v, emb_domain f v, @@ -943,6 +948,11 @@ lemma map_range_neg [neg_zero_class G] [neg_zero_class H] map_range f hf (-v) = -map_range f hf v := ext $ λ _, by simp only [hf', neg_apply, map_range_apply] +lemma map_range_neg' [add_group G] [subtraction_monoid H] [add_monoid_hom_class β G H] + {f : β} (v : α →₀ G) : + map_range f (map_zero f) (-v) = -map_range f (map_zero f) v := +map_range_neg (map_neg f) v + instance [sub_neg_zero_monoid G] : has_sub (α →₀ G) := ⟨zip_with has_sub.sub (sub_zero _)⟩ @[simp] lemma coe_sub [sub_neg_zero_monoid G] (g₁ g₂ : α →₀ G) : ⇑(g₁ - g₂) = g₁ - g₂ := @@ -955,6 +965,11 @@ lemma map_range_sub [sub_neg_zero_monoid G] [sub_neg_zero_monoid H] map_range f hf (v₁ - v₂) = map_range f hf v₁ - map_range f hf v₂ := ext $ λ _, by simp only [hf', sub_apply, map_range_apply] +lemma map_range_sub' [add_group G] [subtraction_monoid H] [add_monoid_hom_class β G H] + {f : β} (v₁ v₂ : α →₀ G) : + map_range f (map_zero f) (v₁ - v₂) = map_range f (map_zero f) v₁ - map_range f (map_zero f) v₂ := +map_range_sub (map_sub f) v₁ v₂ + /-- Note the general `finsupp.has_smul` instance doesn't apply as `ℤ` is not distributive unless `β i`'s addition is commutative. -/ instance has_int_scalar [add_group G] : has_smul ℤ (α →₀ G) := From 8a391fa4fbc14a90441882d413683a20a3d7cf14 Mon Sep 17 00:00:00 2001 From: sgouezel Date: Thu, 5 Jan 2023 13:17:56 +0000 Subject: [PATCH 0145/1291] feat(analysis/convolution): weaken topological assumptions (#18012) We redefine `locally_integrable`: currently, this means "integrable on every compact set", and we change it to "integrable on a neighborhood of any point", to bring it in line with `is_locally_finite_measure`. This makes it possible to weaken a lot of assumptions in the file on convolutions, removing second countability or local compactness assumptions. --- src/analysis/convolution.lean | 164 ++++++++-------- .../function/locally_integrable.lean | 182 +++++++++++++----- .../integral/integrable_on.lean | 32 ++- .../integral/interval_integral.lean | 2 +- src/topology/algebra/group/basic.lean | 26 +++ test/monotonicity.lean | 6 +- 6 files changed, 278 insertions(+), 134 deletions(-) diff --git a/src/analysis/convolution.lean b/src/analysis/convolution.lean index 01acc85d1edd3..cd4b826c51faf 100644 --- a/src/analysis/convolution.lean +++ b/src/analysis/convolution.lean @@ -101,9 +101,9 @@ section no_measurability variables [add_group G] [topological_space G] -lemma has_compact_support.convolution_integrand_bound_right (hcg : has_compact_support g) - (hg : continuous g) {x t : G} {s : set G} (hx : x ∈ s) : - ‖L (f t) (g (x - t))‖ ≤ (- tsupport g + s).indicator (λ t, ‖L‖ * ‖f t‖ * (⨆ i, ‖g i‖)) t := +lemma has_compact_support.convolution_integrand_bound_right_of_subset (hcg : has_compact_support g) + (hg : continuous g) {x t : G} {s u : set G} (hx : x ∈ s) (hu : - tsupport g + s ⊆ u) : + ‖L (f t) (g (x - t))‖ ≤ u.indicator (λ t, ‖L‖ * ‖f t‖ * (⨆ i, ‖g i‖)) t := begin refine le_indicator (λ t ht, _) (λ t ht, _) t, { refine (L.le_op_norm₂ _ _).trans _, @@ -111,11 +111,18 @@ begin (le_csupr (hg.norm.bdd_above_range_of_has_compact_support hcg.norm) $ x - t) (mul_nonneg (norm_nonneg _) (norm_nonneg _)) }, { have : x - t ∉ support g, - { refine mt (λ hxt, _) ht, refine ⟨_, _, set.neg_mem_neg.mpr (subset_closure hxt), hx, _⟩, + { refine mt (λ hxt, _) ht, + apply hu, + refine ⟨_, _, set.neg_mem_neg.mpr (subset_closure hxt), hx, _⟩, rw [neg_sub, sub_add_cancel] }, rw [nmem_support.mp this, (L _).map_zero, norm_zero] } end +lemma has_compact_support.convolution_integrand_bound_right (hcg : has_compact_support g) + (hg : continuous g) {x t : G} {s : set G} (hx : x ∈ s) : + ‖L (f t) (g (x - t))‖ ≤ (- tsupport g + s).indicator (λ t, ‖L‖ * ‖f t‖ * (⨆ i, ‖g i‖)) t := +hcg.convolution_integrand_bound_right_of_subset L hg hx subset.rfl + lemma continuous.convolution_integrand_fst [has_continuous_sub G] (hg : continuous g) (t : G) : continuous (λ x, L (f t) (g (x - t))) := L.continuous₂.comp₂ continuous_const $ hg.comp $ continuous_id.sub continuous_const @@ -156,14 +163,19 @@ variables (L) section group variables [add_group G] -variables [has_measurable_add₂ G] [has_measurable_neg G] -lemma measure_theory.ae_strongly_measurable.convolution_integrand' [sigma_finite ν] +lemma measure_theory.ae_strongly_measurable.convolution_integrand' + [has_measurable_add₂ G] [has_measurable_neg G] [sigma_finite ν] (hf : ae_strongly_measurable f ν) (hg : ae_strongly_measurable g $ map (λ (p : G × G), p.1 - p.2) (μ.prod ν)) : ae_strongly_measurable (λ p : G × G, L (f p.2) (g (p.1 - p.2))) (μ.prod ν) := L.ae_strongly_measurable_comp₂ hf.snd $ hg.comp_measurable measurable_sub + +section + +variables [has_measurable_add G] [has_measurable_neg G] + lemma measure_theory.ae_strongly_measurable.convolution_integrand_snd' (hf : ae_strongly_measurable f μ) {x : G} (hg : ae_strongly_measurable g $ map (λ t, x - t) μ) : @@ -176,21 +188,17 @@ lemma measure_theory.ae_strongly_measurable.convolution_integrand_swap_snd' L.ae_strongly_measurable_comp₂ (hf.comp_measurable $ measurable_id.const_sub x) hg /-- A sufficient condition to prove that `f ⋆[L, μ] g` exists. -We assume that the integrand has compact support and `g` is bounded on this support (note that -both properties hold if `g` is continuous with compact support). We also require that `f` is -integrable on the support of the integrand, and that both functions are strongly measurable. - -Note: we could weaken the measurability condition to hold only for `μ.restrict s`. -/ +We assume that `f` is integrable on a set `s` and `g` is bounded and ae strongly measurable +on `x₀ - s` (note that both properties hold if `g` is continuous with compact support). -/ lemma bdd_above.convolution_exists_at' {x₀ : G} {s : set G} (hbg : bdd_above ((λ i, ‖g i‖) '' ((λ t, - t + x₀) ⁻¹' s))) (hs : measurable_set s) (h2s : support (λ t, L (f t) (g (x₀ - t))) ⊆ s) - (hf : integrable_on f s μ) - (hmf : ae_strongly_measurable f μ) - (hmg : ae_strongly_measurable g $ map (λ t, x₀ - t) μ) : - convolution_exists_at f g x₀ L μ := + (hf : integrable_on f s μ) (hmg : ae_strongly_measurable g $ map (λ t, x₀ - t) (μ.restrict s)) : + convolution_exists_at f g x₀ L μ := begin + rw [convolution_exists_at, ← integrable_on_iff_integrable_of_support_subset h2s hs], set s' := (λ t, - t + x₀) ⁻¹' s, - have : ∀ᵐ (t : G) ∂μ, + have : ∀ᵐ (t : G) ∂(μ.restrict s), ‖L (f t) (g (x₀ - t))‖ ≤ s.indicator (λ t, ‖L‖ * ‖f t‖ * ⨆ i : s', ‖g i‖) t, { refine eventually_of_forall _, refine le_indicator (λ t ht, _) (λ t ht, _), @@ -202,8 +210,8 @@ begin { have : t ∉ support (λ t, L (f t) (g (x₀ - t))) := mt (λ h, h2s h) ht, rw [nmem_support.mp this, norm_zero] } }, refine integrable.mono' _ _ this, - { rw [integrable_indicator_iff hs], exact (hf.norm.const_mul _).mul_const _ }, - { exact hmf.convolution_integrand_snd' L hmg } + { rw [integrable_indicator_iff hs], exact ((hf.norm.const_mul _).mul_const _).integrable_on }, + { exact hf.ae_strongly_measurable.convolution_integrand_snd' L hmg } end /-- If `‖f‖ *[μ] ‖g‖` exists, then `f *[L, μ] g` exists. -/ @@ -219,8 +227,10 @@ begin apply L.le_op_norm₂, end +end + section left -variables [sigma_finite μ] [is_add_right_invariant μ] +variables [has_measurable_add₂ G] [has_measurable_neg G] [sigma_finite μ] [is_add_right_invariant μ] lemma measure_theory.ae_strongly_measurable.convolution_integrand_snd (hf : ae_strongly_measurable f μ) (hg : ae_strongly_measurable g μ) @@ -247,7 +257,8 @@ end left section right -variables [sigma_finite μ] [is_add_right_invariant μ] [sigma_finite ν] +variables [has_measurable_add₂ G] [has_measurable_neg G] +[sigma_finite μ] [is_add_right_invariant μ] [sigma_finite ν] lemma measure_theory.ae_strongly_measurable.convolution_integrand (hf : ae_strongly_measurable f ν) (hg : ae_strongly_measurable g μ) : @@ -283,14 +294,27 @@ lemma measure_theory.integrable.ae_convolution_exists (hf : integrable f ν) (hg end right variables [topological_space G] [topological_add_group G] [borel_space G] - [second_countable_topology G] [sigma_compact_space G] lemma has_compact_support.convolution_exists_at {x₀ : G} (h : has_compact_support (λ t, L (f t) (g (x₀ - t)))) (hf : locally_integrable f μ) (hg : continuous g) : convolution_exists_at f g x₀ L μ := -((((homeomorph.neg G).trans $ homeomorph.add_right x₀).is_compact_preimage.mpr h).bdd_above_image - hg.norm.continuous_on).convolution_exists_at' L is_closed_closure.measurable_set subset_closure - (hf h) hf.ae_strongly_measurable hg.ae_strongly_measurable +begin + let u := (homeomorph.neg G).trans (homeomorph.add_right x₀), + let v := (homeomorph.neg G).trans (homeomorph.add_left x₀), + apply ((u.is_compact_preimage.mpr h).bdd_above_image hg.norm.continuous_on).convolution_exists_at' + L is_closed_closure.measurable_set subset_closure (hf.integrable_on_is_compact h), + have A : ae_strongly_measurable (g ∘ ⇑v) + (μ.restrict (tsupport (λ (t : G), (L (f t)) (g (x₀ - t))))), + { apply (hg.comp v.continuous).continuous_on.ae_strongly_measurable_of_is_compact h, + exact (is_closed_tsupport _).measurable_set }, + convert ((v.continuous.measurable.measure_preserving + (μ.restrict (tsupport (λ t, L (f t) (g (x₀ - t)))))).ae_strongly_measurable_comp_iff + v.to_measurable_equiv.measurable_embedding).1 A, + ext x, + simp only [homeomorph.neg, sub_eq_add_neg, coe_to_add_units, homeomorph.trans_apply, + equiv.neg_apply, equiv.to_fun_as_coe, homeomorph.homeomorph_mk_coe, equiv.coe_fn_mk, + homeomorph.coe_add_left], +end lemma has_compact_support.convolution_exists_right (hcg : has_compact_support g) (hf : locally_integrable f μ) (hg : continuous g) : @@ -322,7 +346,7 @@ variables [add_comm_group G] section measurable_group -variables [has_measurable_add₂ G] [has_measurable_neg G] [is_add_left_invariant μ] +variables [has_measurable_neg G] [is_add_left_invariant μ] /-- A sufficient condition to prove that `f ⋆[L, μ] g` exists. We assume that the integrand has compact support and `g` is bounded on this support (note that @@ -331,21 +355,22 @@ integrable on the support of the integrand, and that both functions are strongly This is a variant of `bdd_above.convolution_exists_at'` in an abelian group with a left-invariant measure. This allows us to state the boundedness and measurability of `g` in a more natural way. -/ -lemma bdd_above.convolution_exists_at [sigma_finite μ] {x₀ : G} +lemma bdd_above.convolution_exists_at [has_measurable_add₂ G] [sigma_finite μ] {x₀ : G} {s : set G} (hbg : bdd_above ((λ i, ‖g i‖) '' ((λ t, x₀ - t) ⁻¹' s))) (hs : measurable_set s) (h2s : support (λ t, L (f t) (g (x₀ - t))) ⊆ s) - (hf : integrable_on f s μ) - (hmf : ae_strongly_measurable f μ) - (hmg : ae_strongly_measurable g μ) : + (hf : integrable_on f s μ) (hmg : ae_strongly_measurable g μ) : convolution_exists_at f g x₀ L μ := begin - refine bdd_above.convolution_exists_at' L _ hs h2s hf hmf _, + refine bdd_above.convolution_exists_at' L _ hs h2s hf _, { simp_rw [← sub_eq_neg_add, hbg] }, - { exact hmg.mono' - (quasi_measure_preserving_sub_left_of_right_invariant μ x₀).absolutely_continuous } + { have : ae_strongly_measurable g (map (λ (t : G), x₀ - t) μ), from hmg.mono' + (quasi_measure_preserving_sub_left_of_right_invariant μ x₀).absolutely_continuous, + apply this.mono_measure, + exact map_mono_of_ae_measurable restrict_le_self + (measurable_const.sub measurable_id').ae_measurable } end -variables {L} [is_neg_invariant μ] +variables {L} [has_measurable_add G] [is_neg_invariant μ] lemma convolution_exists_at_flip : convolution_exists_at g f x L.flip μ ↔ convolution_exists_at f g x L μ := @@ -363,8 +388,7 @@ convolution_exists_at_flip.symm end measurable_group variables [topological_space G] [topological_add_group G] [borel_space G] - [second_countable_topology G] [is_add_left_invariant μ] [is_neg_invariant μ] - [sigma_compact_space G] + [is_add_left_invariant μ] [is_neg_invariant μ] lemma has_compact_support.convolution_exists_left (hcf : has_compact_support f) (hf : continuous f) (hg : locally_integrable g μ) : @@ -489,34 +513,40 @@ is_compact_of_is_closed_subset (hcg.is_compact.add hcf) is_closed_closure $ clos ((support_convolution_subset_swap L).trans $ add_subset_add subset_closure subset_closure) (hcg.is_compact.add hcf).is_closed -variables [borel_space G] [second_countable_topology G] +variables [borel_space G] /-- The convolution is continuous if one function is locally integrable and the other has compact support and is continuous. -/ -lemma has_compact_support.continuous_convolution_right [locally_compact_space G] [t2_space G] +lemma has_compact_support.continuous_convolution_right [first_countable_topology G] (hcg : has_compact_support g) (hf : locally_integrable f μ) (hg : continuous g) : continuous (f ⋆[L, μ] g) := begin refine continuous_iff_continuous_at.mpr (λ x₀, _), - obtain ⟨K, hK, h2K⟩ := exists_compact_mem_nhds x₀, - let K' := - tsupport g + K, - have hK' : is_compact K' := hcg.neg.add hK, + let K' := - tsupport g + {x₀}, + have hK' : is_compact K' := hcg.neg.add is_compact_singleton, + obtain ⟨U, U_open, K'U, hU⟩ : ∃ U, is_open U ∧ K' ⊆ U ∧ integrable_on f U μ, + from hf.integrable_on_nhds_is_compact hK', have : ∀ᶠ x in 𝓝 x₀, ∀ᵐ (t : G) ∂μ, - ‖L (f t) (g (x - t))‖ ≤ K'.indicator (λ t, ‖L‖ * ‖f t‖ * (⨆ i, ‖g i‖)) t := - eventually_of_mem h2K (λ x hx, eventually_of_forall $ - λ t, hcg.convolution_integrand_bound_right L hg hx), + ‖L (f t) (g (x - t))‖ ≤ U.indicator (λ t, ‖L‖ * ‖f t‖ * (⨆ i, ‖g i‖)) t, + { obtain ⟨V, V_mem, hV⟩ : ∃ (V : set G) (H : V ∈ 𝓝 (0 : G)), K' + V ⊆ U, + from compact_open_separated_add_right hK' U_open K'U, + have : {x₀} + V ∈ 𝓝 x₀, from singleton_add_mem_nhds_of_nhds_zero x₀ V_mem, + filter_upwards [this] with x hx, + apply eventually_of_forall (λ t, _), + apply hcg.convolution_integrand_bound_right_of_subset L hg hx, + rwa ← add_assoc }, refine continuous_at_of_dominated _ this _ _, - { exact eventually_of_forall - (λ x, hf.ae_strongly_measurable.convolution_integrand_snd' L hg.ae_strongly_measurable) }, - { rw [integrable_indicator_iff hK'.measurable_set], - exact ((hf hK').norm.const_mul _).mul_const _ }, + { apply eventually_of_forall (λ x, _), + exact (has_compact_support.convolution_exists_right L hcg hf hg x).ae_strongly_measurable }, + { rw [integrable_indicator_iff U_open.measurable_set], + exact (hU.norm.const_mul _).mul_const _ }, { exact eventually_of_forall (λ t, (L.continuous₂.comp₂ continuous_const $ hg.comp $ continuous_id.sub $ by apply continuous_const).continuous_at) } end /-- The convolution is continuous if one function is integrable and the other is bounded and continuous. -/ -lemma bdd_above.continuous_convolution_right_of_integrable +lemma bdd_above.continuous_convolution_right_of_integrable [second_countable_topology G] (hbg : bdd_above (range (λ x, ‖g x‖))) (hf : integrable f μ) (hg : continuous g) : continuous (f ⋆[L, μ] g) := begin @@ -535,14 +565,6 @@ begin hg.comp $ continuous_id.sub $ by apply continuous_const).continuous_at) } end -/-- A version of `has_compact_support.continuous_convolution_right` that works if `G` is -not locally compact but requires that `g` is integrable. -/ -lemma has_compact_support.continuous_convolution_right_of_integrable - (hcg : has_compact_support g) (hf : integrable f μ) (hg : continuous g) : - continuous (f ⋆[L, μ] g) := -(hg.norm.bdd_above_range_of_has_compact_support hcg.norm).continuous_convolution_right_of_integrable - L hf hg - end group section comm_group @@ -600,25 +622,17 @@ end measurable variables [topological_space G] variables [topological_add_group G] variables [borel_space G] -variables [second_countable_topology G] -lemma has_compact_support.continuous_convolution_left [locally_compact_space G] [t2_space G] +lemma has_compact_support.continuous_convolution_left [first_countable_topology G] (hcf : has_compact_support f) (hf : continuous f) (hg : locally_integrable g μ) : - continuous (f ⋆[L, μ] g) := + continuous (f ⋆[L, μ] g) := by { rw [← convolution_flip], exact hcf.continuous_convolution_right L.flip hg hf } -lemma bdd_above.continuous_convolution_left_of_integrable +lemma bdd_above.continuous_convolution_left_of_integrable [second_countable_topology G] (hbf : bdd_above (range (λ x, ‖f x‖))) (hf : continuous f) (hg : integrable g μ) : - continuous (f ⋆[L, μ] g) := + continuous (f ⋆[L, μ] g) := by { rw [← convolution_flip], exact hbf.continuous_convolution_right_of_integrable L.flip hg hf } -/-- A version of `has_compact_support.continuous_convolution_left` that works if `G` is -not locally compact but requires that `g` is integrable. -/ -lemma has_compact_support.continuous_convolution_left_of_integrable - (hcf : has_compact_support f) (hf : continuous f) (hg : integrable g μ) : - continuous (f ⋆[L, μ] g) := -by { rw [← convolution_flip], exact hcf.continuous_convolution_right_of_integrable L.flip hg hf } - end comm_group section normed_add_comm_group @@ -664,7 +678,7 @@ lemma dist_convolution_le' {x₀ : G} {R ε : ℝ} {z₀ : E'} begin have hfg : convolution_exists_at f g x₀ L μ, { refine bdd_above.convolution_exists_at L _ metric.is_open_ball.measurable_set - (subset_trans _ hf) hif.integrable_on hif.ae_strongly_measurable hmg, + (subset_trans _ hf) hif.integrable_on hmg, swap, { refine λ t, mt (λ ht : f t = 0, _), simp_rw [ht, L.map_zero₂] }, rw [bdd_above_def], refine ⟨‖z₀‖ + ε, _⟩, @@ -946,23 +960,18 @@ end end assoc -variables [normed_add_comm_group G] [borel_space G] [normed_space 𝕜 G] +variables [normed_add_comm_group G] [borel_space G] lemma convolution_precompR_apply {g : G → E'' →L[𝕜] E'} (hf : locally_integrable f μ) (hcg : has_compact_support g) (hg : continuous g) (x₀ : G) (x : E'') : (f ⋆[L.precompR E'', μ] g) x₀ x = (f ⋆[L, μ] (λ a, g a x)) x₀ := begin - rcases hcg.eq_zero_or_finite_dimensional 𝕜 hg with rfl|fin_dim, - { simp only [convolution, pi.zero_apply, integral_const, smul_zero, zero_apply, - _root_.map_zero] }, - resetI, - haveI : proper_space G, from finite_dimensional.proper_is_R_or_C 𝕜 G, have := hcg.convolution_exists_right (L.precompR E'' : _) hf hg x₀, simp_rw [convolution_def, continuous_linear_map.integral_apply this], refl, end -variables [sigma_finite μ] [is_add_left_invariant μ] +variables [normed_space 𝕜 G] [sigma_finite μ] [is_add_left_invariant μ] /-- Compute the total derivative of `f ⋆ g` if `g` is `C^1` with compact support and `f` is locally integrable. To write down the total derivative as a convolution, we use @@ -998,7 +1007,7 @@ begin exact (hcg.fderiv 𝕜).convolution_integrand_bound_right L' (hg.continuous_fderiv le_rfl) (ball_subset_closed_ball hx) }, { rw [integrable_indicator_iff hK'.measurable_set], - exact ((hf hK').norm.const_mul _).mul_const _ }, + exact ((hf.integrable_on_is_compact hK').norm.const_mul _).mul_const _ }, { exact eventually_of_forall (λ t x hx, (L _).has_fderiv_at.comp x (h3 x t)) }, end @@ -1017,7 +1026,6 @@ begin rcases hcg.eq_zero_or_finite_dimensional 𝕜 hg.continuous with rfl|fin_dim, { simp only [convolution_zero], exact cont_diff_zero_fun, }, resetI, - haveI : proper_space G, from finite_dimensional.proper_is_R_or_C 𝕜 G, induction n using enat.nat_induction with n ih ih generalizing g, { rw [cont_diff_zero] at hg ⊢, exact hcg.continuous_convolution_right L hf hg }, diff --git a/src/measure_theory/function/locally_integrable.lean b/src/measure_theory/function/locally_integrable.lean index 2deb71c44f857..1e5d2f178bcb2 100644 --- a/src/measure_theory/function/locally_integrable.lean +++ b/src/measure_theory/function/locally_integrable.lean @@ -9,9 +9,9 @@ import measure_theory.integral.integrable_on # Locally integrable functions A function is called *locally integrable* (`measure_theory.locally_integrable`) if it is integrable -on every compact subset of its domain. +on a neighborhood of every point. -This file contains properties of locally integrable functions and of integrability results +This file contains properties of locally integrable functions and integrability results on compact sets. ## Main statements @@ -29,53 +29,99 @@ variables [normed_add_comm_group E] {f : X → E} {μ : measure X} namespace measure_theory -/-- A function `f : X → E` is locally integrable if it is integrable on all compact sets. - See `measure_theory.locally_integrable_iff` for the justification of this name. -/ +/-- A function `f : X → E` is locally integrable if it is integrable on a neighborhood of every +point. In particular, it is integrable on all compact sets, +see `locally_integrable.integrable_on_is_compact`. -/ def locally_integrable (f : X → E) (μ : measure X . volume_tac) : Prop := -∀ ⦃K⦄, is_compact K → integrable_on f K μ +∀ (x : X), integrable_at_filter f (𝓝 x) μ lemma integrable.locally_integrable (hf : integrable f μ) : locally_integrable f μ := -λ K hK, hf.integrable_on +λ x, hf.integrable_at_filter _ -lemma locally_integrable.ae_strongly_measurable [sigma_compact_space X] - (hf : locally_integrable f μ) : - ae_strongly_measurable f μ := +/-- If a function is locally integrable, then it is integrable on an open neighborhood of any +compact set. -/ +lemma locally_integrable.integrable_on_nhds_is_compact (hf : locally_integrable f μ) {k : set X} + (hk : is_compact k) : ∃ u, is_open u ∧ k ⊆ u ∧ integrable_on f u μ := begin - rw [← @restrict_univ _ _ μ, ← Union_compact_covering, ae_strongly_measurable_Union_iff], - exact λ i, (hf $ is_compact_compact_covering X i).ae_strongly_measurable + refine is_compact.induction_on hk _ _ _ _, + { refine ⟨∅, is_open_empty, subset.rfl, integrable_on_empty⟩ }, + { rintros s t hst ⟨u, u_open, tu, hu⟩, + exact ⟨u, u_open, hst.trans tu, hu⟩ }, + { rintros s t ⟨u, u_open, su, hu⟩ ⟨v, v_open, tv, hv⟩, + exact ⟨u ∪ v, u_open.union v_open, union_subset_union su tv, hu.union hv⟩ }, + { assume x hx, + rcases hf x with ⟨u, ux, hu⟩, + rcases mem_nhds_iff.1 ux with ⟨v, vu, v_open, xv⟩, + exact ⟨v, nhds_within_le_nhds (v_open.mem_nhds xv), v, v_open, subset.rfl, hu.mono_set vu⟩ } +end + +/-- If a function is locally integrable, then it is integrable on any compact set. -/ +lemma locally_integrable.integrable_on_is_compact {k : set X} (hf : locally_integrable f μ) + (hk : is_compact k) : integrable_on f k μ := +begin + rcases hf.integrable_on_nhds_is_compact hk with ⟨u, u_open, ku, hu⟩, + exact hu.mono_set ku end lemma locally_integrable_iff [locally_compact_space X] : - locally_integrable f μ ↔ ∀ x : X, ∃ U ∈ 𝓝 x, integrable_on f U μ := + locally_integrable f μ ↔ ∀ (k : set X), is_compact k → integrable_on f k μ := begin - refine ⟨λ hf x, _, λ hf K hK, _⟩, - { obtain ⟨K, hK, h2K⟩ := exists_compact_mem_nhds x, exact ⟨K, h2K, hf hK⟩ }, - { refine is_compact.induction_on hK integrable_on_empty (λ s t hst h, h.mono_set hst) - (λ s t hs ht, integrable_on_union.mpr ⟨hs, ht⟩) (λ x hx, _), - obtain ⟨K, hK, h2K⟩ := hf x, - exact ⟨K, nhds_within_le_nhds hK, h2K⟩ } + refine ⟨λ hf k hk, hf.integrable_on_is_compact hk, λ hf x, _⟩, + obtain ⟨K, hK, h2K⟩ := exists_compact_mem_nhds x, + exact ⟨K, h2K, hf K hK⟩, +end + +lemma locally_integrable.ae_strongly_measurable [second_countable_topology X] + (hf : locally_integrable f μ) : ae_strongly_measurable f μ := +begin + have : ∀ x, ∃ u, is_open u ∧ x ∈ u ∧ integrable_on f u μ, + { assume x, + rcases hf x with ⟨s, hs, h's⟩, + rcases mem_nhds_iff.1 hs with ⟨u, us, u_open, xu⟩, + exact ⟨u, u_open, xu, h's.mono_set us⟩ }, + choose u u_open xu hu using this, + obtain ⟨T, T_count, hT⟩ : ∃ (T : set X), T.countable ∧ (⋃ (i : T), u i) = univ, + { have : (⋃ x, u x) = univ, from eq_univ_of_forall (λ x, mem_Union_of_mem x (xu x)), + rw ← this, + simp only [Union_coe_set, subtype.coe_mk], + exact is_open_Union_countable u u_open }, + haveI : countable T, from countable_coe_iff.mpr T_count, + rw [← @restrict_univ _ _ μ, ← hT, ae_strongly_measurable_Union_iff], + exact λ i, (hu i).ae_strongly_measurable, end lemma locally_integrable_const [is_locally_finite_measure μ] (c : E) : locally_integrable (λ x, c) μ := -λ K hK, by simp only [integrable_on_const, hK.measure_lt_top, or_true] +begin + assume x, + rcases μ.finite_at_nhds x with ⟨U, hU, h'U⟩, + refine ⟨U, hU, _⟩, + simp only [h'U, integrable_on_const, or_true], +end lemma locally_integrable.indicator (hf : locally_integrable f μ) {s : set X} (hs : measurable_set s) : locally_integrable (s.indicator f) μ := -λ K hK, (hf hK).indicator hs +begin + assume x, + rcases hf x with ⟨U, hU, h'U⟩, + exact ⟨U, hU, h'U.indicator hs⟩, +end theorem locally_integrable_map_homeomorph [borel_space X] [borel_space Y] (e : X ≃ₜ Y) {f : Y → E} {μ : measure X} : locally_integrable f (measure.map e μ) ↔ locally_integrable (f ∘ e) μ := begin - refine ⟨λ h k hk, _, λ h k hk, _⟩, - { have : is_compact (e.symm ⁻¹' k), from (homeomorph.is_compact_preimage _).2 hk, - convert (integrable_on_map_equiv e.to_measurable_equiv).1 (h this) using 1, - simp only [←preimage_comp, homeomorph.to_measurable_equiv_coe, homeomorph.symm_comp_self, - preimage_id_eq, id.def] }, - { apply (integrable_on_map_equiv e.to_measurable_equiv).2, - have : is_compact (e ⁻¹' k), from (homeomorph.is_compact_preimage _).2 hk, - exact h this } + refine ⟨λ h x, _, λ h x, _⟩, + { rcases h (e x) with ⟨U, hU, h'U⟩, + refine ⟨e ⁻¹' U, e.continuous.continuous_at.preimage_mem_nhds hU, _⟩, + exact (integrable_on_map_equiv e.to_measurable_equiv).1 h'U }, + { rcases h (e.symm x) with ⟨U, hU, h'U⟩, + refine ⟨e.symm ⁻¹' U, e.symm.continuous.continuous_at.preimage_mem_nhds hU, _⟩, + apply (integrable_on_map_equiv e.to_measurable_equiv).2, + simp only [homeomorph.to_measurable_equiv_coe], + convert h'U, + ext x, + simp only [mem_preimage, homeomorph.symm_apply_apply] } end section mul @@ -138,9 +184,16 @@ is_compact.induction_on hK integrable_on_empty (λ s t hst ht, ht.mono_set hst) section borel -variables [opens_measurable_space X] [metrizable_space X] [is_locally_finite_measure μ] +variables [opens_measurable_space X] [is_locally_finite_measure μ] variables {K : set X} {a b : X} +/-- A continuous function `f` is locally integrable with respect to any locally finite measure. -/ +lemma continuous.locally_integrable [second_countable_topology_either X E] + (hf : continuous f) : locally_integrable f μ := +hf.integrable_at_nhds + +variables [metrizable_space X] + /-- A function `f` continuous on a compact set `K` is integrable on this set with respect to any locally finite measure. -/ lemma continuous_on.integrable_on_compact (hK : is_compact K) (hf : continuous_on f K) : @@ -151,17 +204,13 @@ begin exact hf.integrable_at_nhds_within_of_is_separable hK.measurable_set hK.is_separable hx, end -/-- A continuous function `f` is locally integrable with respect to any locally finite measure. -/ -lemma continuous.locally_integrable (hf : continuous f) : locally_integrable f μ := -λ s hs, hf.continuous_on.integrable_on_compact hs - lemma continuous_on.integrable_on_Icc [preorder X] [compact_Icc_space X] (hf : continuous_on f (Icc a b)) : integrable_on f (Icc a b) μ := hf.integrable_on_compact is_compact_Icc lemma continuous.integrable_on_Icc [preorder X] [compact_Icc_space X] (hf : continuous f) : integrable_on f (Icc a b) μ := -hf.locally_integrable is_compact_Icc +hf.continuous_on.integrable_on_Icc lemma continuous.integrable_on_Ioc [preorder X] [compact_Icc_space X] (hf : continuous f) : integrable_on f (Ioc a b) μ := @@ -182,44 +231,75 @@ hf.integrable_on_Ioc /-- A continuous function with compact support is integrable on the whole space. -/ lemma continuous.integrable_of_has_compact_support (hf : continuous f) (hcf : has_compact_support f) : integrable f μ := -(integrable_on_iff_integable_of_support_subset (subset_tsupport f) measurable_set_closure).mp $ - hf.locally_integrable hcf +(integrable_on_iff_integrable_of_support_subset (subset_tsupport f) measurable_set_closure).mp $ + hf.continuous_on.integrable_on_compact hcf end borel +open_locale ennreal + section monotone -variables [borel_space X] [metrizable_space X] +variables [borel_space X] [conditionally_complete_linear_order X] [conditionally_complete_linear_order E] [order_topology X] [order_topology E] [second_countable_topology E] - [is_locally_finite_measure μ] {s : set X} + {s : set X} -lemma monotone_on.integrable_on_compact (hs : is_compact s) (hmono : monotone_on f s) : +lemma monotone_on.integrable_on_of_measure_ne_top + (hmono : monotone_on f s) {a b : X} (ha : is_least s a) (hb : is_greatest s b) (hs : μ s ≠ ∞) + (h's : measurable_set s) : integrable_on f s μ := begin borelize E, obtain rfl | h := s.eq_empty_or_nonempty, { exact integrable_on_empty }, - have hbelow : bdd_below (f '' s) := - ⟨f (Inf s), λ x ⟨y, hy, hyx⟩, hyx ▸ hmono (hs.Inf_mem h) hy (cInf_le hs.bdd_below hy)⟩, - have habove : bdd_above (f '' s) := - ⟨f (Sup s), λ x ⟨y, hy, hyx⟩, hyx ▸ hmono hy (hs.Sup_mem h) (le_cSup hs.bdd_above hy)⟩, + have hbelow : bdd_below (f '' s) := ⟨f a, λ x ⟨y, hy, hyx⟩, hyx ▸ hmono ha.1 hy (ha.2 hy)⟩, + have habove : bdd_above (f '' s) := ⟨f b, λ x ⟨y, hy, hyx⟩, hyx ▸ hmono hy hb.1 (hb.2 hy)⟩, have : metric.bounded (f '' s) := metric.bounded_of_bdd_above_of_bdd_below habove hbelow, rcases bounded_iff_forall_norm_le.mp this with ⟨C, hC⟩, - refine integrable.mono' (continuous_const.locally_integrable hs) - (ae_measurable_restrict_of_monotone_on hs.measurable_set hmono).ae_strongly_measurable - ((ae_restrict_iff' hs.measurable_set).mpr $ ae_of_all _ $ + have A : integrable_on (λ x, C) s μ, by simp only [hs.lt_top, integrable_on_const, or_true], + refine integrable.mono' A + (ae_measurable_restrict_of_monotone_on h's hmono).ae_strongly_measurable + ((ae_restrict_iff' h's).mpr $ ae_of_all _ $ λ y hy, hC (f y) (mem_image_of_mem f hy)), end -lemma antitone_on.integrable_on_compact (hs : is_compact s) (hanti : antitone_on f s) : +lemma monotone_on.integrable_on_is_compact [is_finite_measure_on_compacts μ] + (hs : is_compact s) (hmono : monotone_on f s) : integrable_on f s μ := -hanti.dual_right.integrable_on_compact hs +begin + obtain rfl | h := s.eq_empty_or_nonempty, + { exact integrable_on_empty }, + { exact hmono.integrable_on_of_measure_ne_top (hs.is_least_Inf h) (hs.is_greatest_Sup h) + hs.measure_lt_top.ne hs.measurable_set } +end -lemma monotone.locally_integrable (hmono : monotone f) : locally_integrable f μ := -λ s hs, (hmono.monotone_on _).integrable_on_compact hs +lemma antitone_on.integrable_on_of_measure_ne_top + (hanti : antitone_on f s) {a b : X} (ha : is_least s a) (hb : is_greatest s b) (hs : μ s ≠ ∞) + (h's : measurable_set s) : + integrable_on f s μ := +hanti.dual_right.integrable_on_of_measure_ne_top ha hb hs h's + +lemma antione_on.integrable_on_is_compact [is_finite_measure_on_compacts μ] + (hs : is_compact s) (hanti : antitone_on f s) : + integrable_on f s μ := +hanti.dual_right.integrable_on_is_compact hs + +lemma monotone.locally_integrable [is_locally_finite_measure μ] (hmono : monotone f) : + locally_integrable f μ := +begin + assume x, + rcases μ.finite_at_nhds x with ⟨U, hU, h'U⟩, + obtain ⟨a, b, xab, hab, abU⟩ : ∃ (a b : X), x ∈ Icc a b ∧ Icc a b ∈ 𝓝 x ∧ Icc a b ⊆ U, + from exists_Icc_mem_subset_of_mem_nhds hU, + have ab : a ≤ b := xab.1.trans xab.2, + refine ⟨Icc a b, hab, _⟩, + exact (hmono.monotone_on _).integrable_on_of_measure_ne_top (is_least_Icc ab) + (is_greatest_Icc ab) ((measure_mono abU).trans_lt h'U).ne measurable_set_Icc, +end -lemma antitone.locally_integrable (hanti : antitone f) : locally_integrable f μ := +lemma antitone.locally_integrable [is_locally_finite_measure μ] (hanti : antitone f) : + locally_integrable f μ := hanti.dual_right.locally_integrable end monotone diff --git a/src/measure_theory/integral/integrable_on.lean b/src/measure_theory/integral/integrable_on.lean index dca0a9e1719a1..95432b32f04f5 100644 --- a/src/measure_theory/integral/integrable_on.lean +++ b/src/measure_theory/integral/integrable_on.lean @@ -231,7 +231,7 @@ begin simpa only [set.univ_inter, measurable_set.univ, measure.restrict_apply] using hμs, end -lemma integrable_on_iff_integable_of_support_subset {f : α → E} {s : set α} +lemma integrable_on_iff_integrable_of_support_subset {f : α → E} {s : set α} (h1s : support f ⊆ s) (h2s : measurable_set s) : integrable_on f s μ ↔ integrable f μ := begin @@ -267,6 +267,10 @@ def integrable_at_filter (f : α → E) (l : filter α) (μ : measure α . volum variables {l l' : filter α} +lemma integrable.integrable_at_filter (h : integrable f μ) (l : filter α) : + integrable_at_filter f l μ := +⟨univ, filter.univ_mem, integrable_on_univ.2 h⟩ + protected lemma integrable_at_filter.eventually (h : integrable_at_filter f l μ) : ∀ᶠ s in l.small_sets, integrable_on f s μ := iff.mpr (eventually_small_sets' $ λ s t hst ht, ht.mono_set hst) h @@ -405,6 +409,22 @@ begin { exact is_separable_of_separable_space _ } end +/-- A function which is continuous on a compact set `s` is almost everywhere strongly measurable +with respect to `μ.restrict s`. -/ +lemma continuous_on.ae_strongly_measurable_of_is_compact + [topological_space α] [opens_measurable_space α] [topological_space β] [pseudo_metrizable_space β] + {f : α → β} {s : set α} {μ : measure α} + (hf : continuous_on f s) (hs : is_compact s) (h's : measurable_set s) : + ae_strongly_measurable f (μ.restrict s) := +begin + letI := pseudo_metrizable_space_pseudo_metric β, + borelize β, + rw ae_strongly_measurable_iff_ae_measurable_separable, + refine ⟨hf.ae_measurable h's, f '' s, _, _⟩, + { exact (hs.image_of_continuous_on hf).is_separable }, + { exact mem_of_superset (self_mem_ae_restrict h's) (subset_preimage_image _ _) } +end + lemma continuous_on.integrable_at_nhds_within_of_is_separable [topological_space α] [pseudo_metrizable_space α] [opens_measurable_space α] {μ : measure α} [is_locally_finite_measure μ] @@ -428,6 +448,16 @@ begin (μ.finite_at_nhds_within _ _), end +lemma continuous.integrable_at_nhds + [topological_space α] [second_countable_topology_either α E] + [opens_measurable_space α] {μ : measure α} [is_locally_finite_measure μ] + {f : α → E} (hf : continuous f) (a : α) : + integrable_at_filter f (𝓝 a) μ := +begin + rw ← nhds_within_univ, + exact hf.continuous_on.integrable_at_nhds_within measurable_set.univ (mem_univ a), +end + /-- If a function is continuous on an open set `s`, then it is strongly measurable at the filter `𝓝 x` for all `x ∈ s` if either the source space or the target space is second-countable. -/ lemma continuous_on.strongly_measurable_at_filter [topological_space α] diff --git a/src/measure_theory/integral/interval_integral.lean b/src/measure_theory/integral/interval_integral.lean index 2f18a52d68e12..20f10a469b7fd 100644 --- a/src/measure_theory/integral/interval_integral.lean +++ b/src/measure_theory/integral/interval_integral.lean @@ -459,7 +459,7 @@ lemma monotone_on.interval_integrable {u : ℝ → E} {a b : ℝ} (hu : monotone interval_integrable u μ a b := begin rw interval_integrable_iff, - exact (hu.integrable_on_compact is_compact_interval).mono_set Ioc_subset_Icc_self, + exact (hu.integrable_on_is_compact is_compact_interval).mono_set Ioc_subset_Icc_self, end lemma antitone_on.interval_integrable {u : ℝ → E} {a b : ℝ} (hu : antitone_on u (interval a b)) : diff --git a/src/topology/algebra/group/basic.lean b/src/topology/algebra/group/basic.lean index 24df65abf3824..f37ec1565a663 100644 --- a/src/topology/algebra/group/basic.lean +++ b/src/topology/algebra/group/basic.lean @@ -955,6 +955,13 @@ by { rw ←bUnion_smul_set, exact is_open_bUnion (λ a _, ht.smul _) } @[to_additive] lemma subset_interior_smul_right : s • interior t ⊆ interior (s • t) := interior_maximal (set.smul_subset_smul_left interior_subset) is_open_interior.smul_left +@[to_additive] lemma smul_mem_nhds (a : α) {x : β} (ht : t ∈ 𝓝 x) : + a • t ∈ 𝓝 (a • x) := +begin + rcases mem_nhds_iff.1 ht with ⟨u, ut, u_open, hu⟩, + exact mem_nhds_iff.2 ⟨a • u, smul_set_mono ut, u_open.smul a, smul_mem_smul_set hu⟩, +end + variables [topological_space α] @[to_additive] lemma subset_interior_smul : interior s • interior t ⊆ interior (s • t) := @@ -973,6 +980,14 @@ subset_interior_smul_right @[to_additive] lemma subset_interior_mul : interior s * interior t ⊆ interior (s * t) := subset_interior_smul +@[to_additive] lemma singleton_mul_mem_nhds (a : α) {b : α} (h : s ∈ 𝓝 b) : + {a} * s ∈ 𝓝 (a * b) := +by { have := smul_mem_nhds a h, rwa ← singleton_smul at this } + +@[to_additive] lemma singleton_mul_mem_nhds_of_nhds_one (a : α) (h : s ∈ 𝓝 (1 : α)) : + {a} * s ∈ 𝓝 a := +by simpa only [mul_one] using singleton_mul_mem_nhds a h + end has_continuous_const_smul section has_continuous_const_smul_op @@ -987,6 +1002,17 @@ interior_maximal (set.mul_subset_mul_right interior_subset) is_open_interior.mul @[to_additive] lemma subset_interior_mul' : interior s * interior t ⊆ interior (s * t) := (set.mul_subset_mul_left interior_subset).trans subset_interior_mul_left +@[to_additive] lemma mul_singleton_mem_nhds (a : α) {b : α} (h : s ∈ 𝓝 b) : + s * {a} ∈ 𝓝 (b * a) := +begin + simp only [←bUnion_op_smul_set, mem_singleton_iff, Union_Union_eq_left], + exact smul_mem_nhds _ h, +end + +@[to_additive] lemma mul_singleton_mem_nhds_of_nhds_one (a : α) (h : s ∈ 𝓝 (1 : α)) : + s * {a} ∈ 𝓝 a := +by simpa only [one_mul] using mul_singleton_mem_nhds a h + end has_continuous_const_smul_op section topological_group diff --git a/test/monotonicity.lean b/test/monotonicity.lean index 30ee14d29d5ca..c4ab5d00dbac9 100644 --- a/test/monotonicity.lean +++ b/test/monotonicity.lean @@ -442,9 +442,9 @@ end example : ∫ x in Icc 0 1, real.exp x ≤ ∫ x in Icc 0 1, real.exp (x+1) := begin mono, - { exact real.continuous_exp.locally_integrable is_compact_Icc }, - { exact (real.continuous_exp.comp $ continuous_add_right 1).locally_integrable - is_compact_Icc }, + { exact real.continuous_exp.locally_integrable.integrable_on_is_compact is_compact_Icc }, + { exact (real.continuous_exp.comp $ continuous_add_right 1) + .locally_integrable.integrable_on_is_compact is_compact_Icc }, intro x, dsimp only, mono, From 50088c97b6065dbe5399e9fe0a0c6f9b6d6db848 Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Thu, 5 Jan 2023 13:17:57 +0000 Subject: [PATCH 0146/1291] feat(ring_theory/dedekind_domain/ideal): add more span_singleton lemmas (#18047) --- src/ring_theory/dedekind_domain/ideal.lean | 26 +++++++++++++++++++--- 1 file changed, 23 insertions(+), 3 deletions(-) diff --git a/src/ring_theory/dedekind_domain/ideal.lean b/src/ring_theory/dedekind_domain/ideal.lean index 6ac888ca97883..157f23c61278e 100644 --- a/src/ring_theory/dedekind_domain/ideal.lean +++ b/src/ring_theory/dedekind_domain/ideal.lean @@ -129,11 +129,31 @@ open submodule submodule.is_principal @[simp] lemma span_singleton_inv (x : K) : (span_singleton R₁⁰ x)⁻¹ = span_singleton _ x⁻¹ := one_div_span_singleton x +@[simp] lemma span_singleton_div_span_singleton (x y : K) : + span_singleton R₁⁰ x / span_singleton R₁⁰ y = span_singleton R₁⁰ (x / y) := +by rw [div_span_singleton, mul_comm, span_singleton_mul_span_singleton, div_eq_mul_inv] + +lemma span_singleton_div_self {x : K} (hx : x ≠ 0) : + span_singleton R₁⁰ x / span_singleton R₁⁰ x = 1 := +by rw [span_singleton_div_span_singleton, div_self hx, span_singleton_one] + +lemma coe_ideal_span_singleton_div_self {x : R₁} (hx : x ≠ 0) : + ((ideal.span {x} : ideal R₁) : fractional_ideal R₁⁰ K) / (ideal.span {x} : ideal R₁) = 1 := +by rw [coe_ideal_span_singleton, span_singleton_div_self K $ + (map_ne_zero_iff _ $ no_zero_smul_divisors.algebra_map_injective R₁ K).mpr hx] + +lemma span_singleton_mul_inv {x : K} (hx : x ≠ 0) : + span_singleton R₁⁰ x * (span_singleton R₁⁰ x)⁻¹ = 1 := +by rw [span_singleton_inv, span_singleton_mul_span_singleton, mul_inv_cancel hx, span_singleton_one] + lemma coe_ideal_span_singleton_mul_inv {x : R₁} (hx : x ≠ 0) : ((ideal.span {x} : ideal R₁) : fractional_ideal R₁⁰ K) * (ideal.span {x} : ideal R₁)⁻¹ = 1 := -by rw [coe_ideal_span_singleton, span_singleton_inv, span_singleton_mul_span_singleton, - mul_inv_cancel $ (map_ne_zero_iff _ $ no_zero_smul_divisors.algebra_map_injective R₁ K).mpr - hx, span_singleton_one] +by rw [coe_ideal_span_singleton, span_singleton_mul_inv K $ + (map_ne_zero_iff _ $ no_zero_smul_divisors.algebra_map_injective R₁ K).mpr hx] + +lemma span_singleton_inv_mul {x : K} (hx : x ≠ 0) : + (span_singleton R₁⁰ x)⁻¹ * span_singleton R₁⁰ x = 1 := +by rw [mul_comm, span_singleton_mul_inv K hx] lemma coe_ideal_span_singleton_inv_mul {x : R₁} (hx : x ≠ 0) : ((ideal.span {x} : ideal R₁) : fractional_ideal R₁⁰ K)⁻¹ * (ideal.span {x} : ideal R₁) = 1 := From baba818b9acea366489e8ba32d2cc0fcaf50a1f7 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Thu, 5 Jan 2023 13:17:58 +0000 Subject: [PATCH 0147/1291] chore(*): add mathlib4 synchronization comments (#18057) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.gcd_monoid.basic` * `data.bundle` * `data.pnat.prime` * `data.rat.cast` * `data.set.intervals.monotone` * `group_theory.group_action.embedding` * `group_theory.submonoid.center` * `group_theory.submonoid.centralizer` * `group_theory.submonoid.operations` * `group_theory.subsemigroup.membership` * `order.extension.linear` * `order.monotone.monovary` * `order.succ_pred.limit` --- src/algebra/gcd_monoid/basic.lean | 3 +++ src/data/bundle.lean | 3 +++ src/data/pnat/prime.lean | 3 +++ src/data/rat/cast.lean | 3 +++ src/data/set/intervals/monotone.lean | 3 +++ src/group_theory/group_action/embedding.lean | 3 +++ src/group_theory/submonoid/center.lean | 3 +++ src/group_theory/submonoid/centralizer.lean | 3 +++ src/group_theory/submonoid/operations.lean | 3 +++ src/group_theory/subsemigroup/membership.lean | 3 +++ src/order/extension/linear.lean | 3 +++ src/order/monotone/monovary.lean | 3 +++ src/order/succ_pred/limit.lean | 3 +++ 13 files changed, 39 insertions(+) diff --git a/src/algebra/gcd_monoid/basic.lean b/src/algebra/gcd_monoid/basic.lean index bce40f5f38246..8a98908d5908f 100644 --- a/src/algebra/gcd_monoid/basic.lean +++ b/src/algebra/gcd_monoid/basic.lean @@ -11,6 +11,9 @@ import algebra.ring.regular /-! # Monoids with normalization functions, `gcd`, and `lcm` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines extra structures on `cancel_comm_monoid_with_zero`s, including `is_domain`s. ## Main Definitions diff --git a/src/data/bundle.lean b/src/data/bundle.lean index 4b05142e2f0a4..68cfd8bbb915b 100644 --- a/src/data/bundle.lean +++ b/src/data/bundle.lean @@ -7,6 +7,9 @@ import algebra.module.basic /-! # Bundle + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. Basic data structure to implement fiber bundles, vector bundles (maybe fibrations?), etc. This file should contain all possible results that do not involve any topology. diff --git a/src/data/pnat/prime.lean b/src/data/pnat/prime.lean index 6427d339b8c0a..4db050a84b83c 100644 --- a/src/data/pnat/prime.lean +++ b/src/data/pnat/prime.lean @@ -9,6 +9,9 @@ import data.pnat.basic /-! # Primality and GCD on pnat +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file extends the theory of `ℕ+` with `gcd`, `lcm` and `prime` functions, analogous to those on `nat`. -/ diff --git a/src/data/rat/cast.lean b/src/data/rat/cast.lean index eba360aac6d30..50bd3848cea21 100644 --- a/src/data/rat/cast.lean +++ b/src/data/rat/cast.lean @@ -13,6 +13,9 @@ import algebra.order.field.basic /-! # Casts for Rational Numbers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Summary We define the canonical injection from ℚ into an arbitrary division ring and prove various diff --git a/src/data/set/intervals/monotone.lean b/src/data/set/intervals/monotone.lean index c9ce045e52930..90f52eb82e83a 100644 --- a/src/data/set/intervals/monotone.lean +++ b/src/data/set/intervals/monotone.lean @@ -9,6 +9,9 @@ import order.succ_pred.basic /-! # Monotonicity on intervals +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove that `set.Ici` etc are monotone/antitone functions. We also prove some lemmas about functions monotone on intervals in `succ_order`s. -/ diff --git a/src/group_theory/group_action/embedding.lean b/src/group_theory/group_action/embedding.lean index bce58765350a1..459fc50b60e86 100644 --- a/src/group_theory/group_action/embedding.lean +++ b/src/group_theory/group_action/embedding.lean @@ -9,6 +9,9 @@ import group_theory.group_action.pi /-! # Group actions on embeddings +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file provides a `mul_action G (α ↪ β)` instance that agrees with the `mul_action G (α → β)` instances defined by `pi.mul_action`. diff --git a/src/group_theory/submonoid/center.lean b/src/group_theory/submonoid/center.lean index c41a592df9157..5e101f7bb870f 100644 --- a/src/group_theory/submonoid/center.lean +++ b/src/group_theory/submonoid/center.lean @@ -9,6 +9,9 @@ import group_theory.subsemigroup.center /-! # Centers of monoids +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions * `submonoid.center`: the center of a monoid diff --git a/src/group_theory/submonoid/centralizer.lean b/src/group_theory/submonoid/centralizer.lean index ce558ee163312..f05395a009469 100644 --- a/src/group_theory/submonoid/centralizer.lean +++ b/src/group_theory/submonoid/centralizer.lean @@ -9,6 +9,9 @@ import group_theory.submonoid.center /-! # Centralizers of magmas and monoids +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions * `submonoid.centralizer`: the centralizer of a subset of a monoid diff --git a/src/group_theory/submonoid/operations.lean b/src/group_theory/submonoid/operations.lean index ab197c4f79197..6dcaf235d28e3 100644 --- a/src/group_theory/submonoid/operations.lean +++ b/src/group_theory/submonoid/operations.lean @@ -12,6 +12,9 @@ import group_theory.subsemigroup.operations /-! # Operations on `submonoid`s +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define various operations on `submonoid`s and `monoid_hom`s. ## Main definitions diff --git a/src/group_theory/subsemigroup/membership.lean b/src/group_theory/subsemigroup/membership.lean index 4a8ded253af3f..90d4d0e144b89 100644 --- a/src/group_theory/subsemigroup/membership.lean +++ b/src/group_theory/subsemigroup/membership.lean @@ -8,6 +8,9 @@ import group_theory.subsemigroup.basic /-! # Subsemigroups: membership criteria +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove various facts about membership in a subsemigroup. The intent is to mimic `group_theory/submonoid/membership`, but currently this file is mostly a stub and only provides rudimentary support. diff --git a/src/order/extension/linear.lean b/src/order/extension/linear.lean index cd8d7fc19e2f9..cdfb97de16c56 100644 --- a/src/order/extension/linear.lean +++ b/src/order/extension/linear.lean @@ -9,6 +9,9 @@ import tactic.by_contra /-! # Extend a partial order to a linear order +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file constructs a linear order which is an extension of the given partial order, using Zorn's lemma. -/ diff --git a/src/order/monotone/monovary.lean b/src/order/monotone/monovary.lean index d2ff137b04331..d050976376e41 100644 --- a/src/order/monotone/monovary.lean +++ b/src/order/monotone/monovary.lean @@ -8,6 +8,9 @@ import data.set.image /-! # Monovariance of functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Two functions *vary together* if a strict change in the first implies a change in the second. This is in some sense a way to say that two functions `f : ι → α`, `g : ι → β` are "monotone diff --git a/src/order/succ_pred/limit.lean b/src/order/succ_pred/limit.lean index cdd69c0596b33..295cf5ed677a1 100644 --- a/src/order/succ_pred/limit.lean +++ b/src/order/succ_pred/limit.lean @@ -9,6 +9,9 @@ import order.succ_pred.basic /-! # Successor and predecessor limits +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define the predicate `order.is_succ_limit` for "successor limits", values that don't cover any others. They are so named since they can't be the successors of anything smaller. We define `order.is_pred_limit` analogously, and prove basic results. From 8f9b933bf5a871305c8ae3485ff43f2e0bb16cfc Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Thu, 5 Jan 2023 13:17:59 +0000 Subject: [PATCH 0148/1291] feat(probability/kernel/basic): deterministic and constant kernels, restriction of a kernel (#18060) --- src/probability/kernel/basic.lean | 113 +++++++++++++++++++++++++++++- 1 file changed, 112 insertions(+), 1 deletion(-) diff --git a/src/probability/kernel/basic.lean b/src/probability/kernel/basic.lean index 37c9c54684dc2..48a1852a37ef7 100644 --- a/src/probability/kernel/basic.lean +++ b/src/probability/kernel/basic.lean @@ -12,7 +12,7 @@ import measure_theory.constructions.prod A kernel from a measurable space `α` to another measurable space `β` is a measurable map `α → measure β`, where the measurable space instance on `measure β` is the one defined in `measure_theory.measure.measurable_space`. That is, a kernel `κ` verifies that for all measurable -sets `s` of `β`, `λ a, κ a s` is measurable. +sets `s` of `β`, `a ↦ κ a s` is measurable. ## Main definitions @@ -289,6 +289,117 @@ end end s_finite +section deterministic + +/-- Kernel which to `a` associates the dirac measure at `f a`. This is a Markov kernel. -/ +noncomputable +def deterministic {f : α → β} (hf : measurable f) : + kernel α β := +{ val := λ a, measure.dirac (f a), + property := + begin + refine measure.measurable_of_measurable_coe _ (λ s hs, _), + simp_rw measure.dirac_apply' _ hs, + refine measurable.indicator _ (hf hs), + simp only [pi.one_apply, measurable_const], + end, } + +lemma deterministic_apply {f : α → β} (hf : measurable f) (a : α) : + deterministic hf a = measure.dirac (f a) := rfl + +lemma deterministic_apply' {f : α → β} (hf : measurable f) (a : α) {s : set β} + (hs : measurable_set s) : + deterministic hf a s = s.indicator (λ _, 1) (f a) := +begin + rw [deterministic], + change measure.dirac (f a) s = s.indicator 1 (f a), + simp_rw measure.dirac_apply' _ hs, +end + +instance is_markov_kernel_deterministic {f : α → β} (hf : measurable f) : + is_markov_kernel (deterministic hf) := +⟨λ a, by { rw deterministic_apply hf, apply_instance, }⟩ + +end deterministic + +section const + +omit mα mβ + +/-- Constant kernel, which always returns the same measure. -/ +def const (α : Type*) {β : Type*} [measurable_space α] {mβ : measurable_space β} (μβ : measure β) : + kernel α β := +{ val := λ _, μβ, + property := measure.measurable_of_measurable_coe _ (λ s hs, measurable_const), } + +include mα mβ + +instance is_finite_kernel_const {μβ : measure β} [hμβ : is_finite_measure μβ] : + is_finite_kernel (const α μβ) := +⟨⟨μβ set.univ, measure_lt_top _ _, λ a, le_rfl⟩⟩ + +instance is_markov_kernel_const {μβ : measure β} [hμβ : is_probability_measure μβ] : + is_markov_kernel (const α μβ) := +⟨λ a, hμβ⟩ + +end const + +omit mα + +/-- In a countable space with measurable singletons, every function `α → measure β` defines a +kernel. -/ +def of_fun_of_countable [measurable_space α] {mβ : measurable_space β} + [countable α] [measurable_singleton_class α] (f : α → measure β) : + kernel α β := +{ val := f, + property := measurable_of_countable f } + +include mα + +section restrict +variables {s t : set β} + +/-- Kernel given by the restriction of the measures in the image of a kernel to a set. -/ +protected noncomputable +def restrict (κ : kernel α β) (hs : measurable_set s) : kernel α β := +{ val := λ a, (κ a).restrict s, + property := + begin + refine measure.measurable_of_measurable_coe _ (λ t ht, _), + simp_rw measure.restrict_apply ht, + exact kernel.measurable_coe κ (ht.inter hs), + end, } + +lemma restrict_apply (κ : kernel α β) (hs : measurable_set s) (a : α) : + kernel.restrict κ hs a = (κ a).restrict s := rfl + +lemma restrict_apply' (κ : kernel α β) (hs : measurable_set s) (a : α) (ht : measurable_set t) : + kernel.restrict κ hs a t = (κ a) (t ∩ s) := +by rw [restrict_apply κ hs a, measure.restrict_apply ht] + +lemma lintegral_restrict (κ : kernel α β) (hs : measurable_set s) (a : α) (f : β → ℝ≥0∞) : + ∫⁻ b, f b ∂(kernel.restrict κ hs a) = ∫⁻ b in s, f b ∂(κ a) := +by rw restrict_apply + +instance is_finite_kernel.restrict (κ : kernel α β) [is_finite_kernel κ] (hs : measurable_set s) : + is_finite_kernel (kernel.restrict κ hs) := +begin + refine ⟨⟨is_finite_kernel.bound κ, is_finite_kernel.bound_lt_top κ, λ a, _⟩⟩, + rw restrict_apply' κ hs a measurable_set.univ, + exact measure_le_bound κ a _, +end + +instance is_s_finite_kernel.restrict (κ : kernel α β) [is_s_finite_kernel κ] + (hs : measurable_set s) : + is_s_finite_kernel (kernel.restrict κ hs) := +begin + refine ⟨⟨λ n, kernel.restrict (seq κ n) hs, infer_instance, _⟩⟩, + ext1 a, + simp_rw [sum_apply, restrict_apply, ← measure.restrict_sum _ hs, ← sum_apply, kernel_sum_seq], +end + +end restrict + end kernel end probability_theory From ed005a8c91df8f75bf3a92255114659c2dad9907 Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Thu, 5 Jan 2023 13:18:01 +0000 Subject: [PATCH 0149/1291] feat(topology/order/basic): replace `partial_order` by `preorder` (#18064) It turns out nearly all `partial_order` hypotheses can be generalized to `preorder` in this file. --- src/topology/order/basic.lean | 52 +++++++++++++++++------------------ 1 file changed, 26 insertions(+), 26 deletions(-) diff --git a/src/topology/order/basic.lean b/src/topology/order/basic.lean index 460b8dd82d41c..3e2512eff625d 100644 --- a/src/topology/order/basic.lean +++ b/src/topology/order/basic.lean @@ -87,7 +87,7 @@ set of points `(x, y)` with `x ≤ y` is closed in the product space. We introdu This property is satisfied for the order topology on a linear order, but it can be satisfied more generally, and suffices to derive many interesting properties relating order and topology. -/ class order_closed_topology (α : Type*) [topological_space α] [preorder α] : Prop := -(is_closed_le' : is_closed {p:α×α | p.1 ≤ p.2}) +(is_closed_le' : is_closed {p : α × α | p.1 ≤ p.2}) instance [topological_space α] [h : first_countable_topology α] : first_countable_topology αᵒᵈ := h @@ -718,7 +718,7 @@ it on a preorder, but it is mostly interesting in linear orders, where it is als We define it as a mixin. If you want to introduce the order topology on a preorder, use `preorder.topology`. -/ class order_topology (α : Type*) [t : topological_space α] [preorder α] : Prop := -(topology_eq_generate_intervals : t = generate_from {s | ∃a, s = Ioi a ∨ s = Iio a}) +(topology_eq_generate_intervals : t = generate_from {s | ∃ a, s = Ioi a ∨ s = Iio a}) /-- (Order) topology on a partial order `α` generated by the subbase of open intervals `(a, ∞) = { x ∣ a < x }, (-∞ , b) = {x ∣ x < b}` for all `a, b` in `α`. We do not register it as an @@ -729,23 +729,23 @@ generate_from {s : set α | ∃ (a : α), s = {b : α | a < b} ∨ s = {b : α | section order_topology -instance {α : Type*} [topological_space α] [partial_order α] [order_topology α] : - order_topology αᵒᵈ := -⟨by convert @order_topology.topology_eq_generate_intervals α _ _ _; - conv in (_ ∨ _) { rw or.comm }; refl⟩ +section preorder -section partial_order -variables [topological_space α] [partial_order α] [t : order_topology α] +variables [topological_space α] [preorder α] [t : order_topology α] include t +instance : order_topology αᵒᵈ := +⟨by convert @order_topology.topology_eq_generate_intervals α _ _ _; + conv in (_ ∨ _) { rw or.comm }; refl⟩ + lemma is_open_iff_generate_intervals {s : set α} : - is_open s ↔ generate_open {s | ∃a, s = Ioi a ∨ s = Iio a} s := + is_open s ↔ generate_open {s | ∃ a, s = Ioi a ∨ s = Iio a} s := by rw [t.topology_eq_generate_intervals]; refl -lemma is_open_lt' (a : α) : is_open {b:α | a < b} := +lemma is_open_lt' (a : α) : is_open {b : α | a < b} := by rw [@is_open_iff_generate_intervals α _ _ t]; exact generate_open.basic _ ⟨a, or.inl rfl⟩ -lemma is_open_gt' (a : α) : is_open {b:α | b < a} := +lemma is_open_gt' (a : α) : is_open {b : α | b < a} := by rw [@is_open_iff_generate_intervals α _ _ t]; exact generate_open.basic _ ⟨a, or.inr rfl⟩ lemma lt_mem_nhds {a b : α} (h : a < b) : ∀ᶠ x in 𝓝 b, a < x := @@ -761,7 +761,7 @@ lemma ge_mem_nhds {a b : α} (h : a < b) : ∀ᶠ x in 𝓝 a, x ≤ b := (𝓝 a).sets_of_superset (gt_mem_nhds h) $ assume b hb, le_of_lt hb lemma nhds_eq_order (a : α) : - 𝓝 a = (⨅b ∈ Iio a, 𝓟 (Ioi b)) ⊓ (⨅b ∈ Ioi a, 𝓟 (Iio b)) := + 𝓝 a = (⨅ b ∈ Iio a, 𝓟 (Ioi b)) ⊓ (⨅ b ∈ Ioi a, 𝓟 (Iio b)) := by rw [t.topology_eq_generate_intervals, nhds_generate_from]; from le_antisymm (le_inf @@ -823,7 +823,7 @@ by rw [nhds_order_unbounded hu hl]; from (tendsto_infi.2 $ assume l, tendsto_infi.2 $ assume hl, tendsto_infi.2 $ assume u, tendsto_infi.2 $ assume hu, tendsto_principal.2 $ h l u hl hu) -end partial_order +end preorder instance tendsto_Ixx_nhds_within {α : Type*} [preorder α] [topological_space α] (a : α) {s t : set α} {Ixx} @@ -832,7 +832,7 @@ instance tendsto_Ixx_nhds_within {α : Type*} [preorder α] [topological_space filter.tendsto_Ixx_class_inf instance tendsto_Icc_class_nhds_pi {ι : Type*} {α : ι → Type*} - [Π i, partial_order (α i)] [Π i, topological_space (α i)] [∀ i, order_topology (α i)] + [Π i, preorder (α i)] [Π i, topological_space (α i)] [∀ i, order_topology (α i)] (f : Π i, α i) : tendsto_Ixx_class Icc (𝓝 f) (𝓝 f) := begin @@ -846,7 +846,7 @@ begin end theorem induced_order_topology' {α : Type u} {β : Type v} - [partial_order α] [ta : topological_space β] [partial_order β] [order_topology β] + [preorder α] [ta : topological_space β] [preorder β] [order_topology β] (f : α → β) (hf : ∀ {x y}, f x < f y ↔ x < y) (H₁ : ∀ {a x}, x < f a → ∃ b < a, x ≤ f b) (H₂ : ∀ {a x}, f a < x → ∃ b > a, f b ≤ x) : @@ -875,7 +875,7 @@ begin end theorem induced_order_topology {α : Type u} {β : Type v} - [partial_order α] [ta : topological_space β] [partial_order β] [order_topology β] + [preorder α] [ta : topological_space β] [preorder β] [order_topology β] (f : α → β) (hf : ∀ {x y}, f x < f y ↔ x < y) (H : ∀ {x y}, x < y → ∃ a, x < f a ∧ f a < y) : @order_topology _ (induced f ta) _ := @@ -932,7 +932,7 @@ begin exact λ hx, ht.out a.2 y.2 ⟨le_of_lt h, le_of_not_gt hx⟩ } } end -lemma nhds_within_Ici_eq'' [topological_space α] [partial_order α] [order_topology α] (a : α) : +lemma nhds_within_Ici_eq'' [topological_space α] [preorder α] [order_topology α] (a : α) : 𝓝[≥] a = (⨅ u (hu : a < u), 𝓟 (Iio u)) ⊓ 𝓟 (Ici a) := begin rw [nhds_within, nhds_eq_order], @@ -940,16 +940,16 @@ begin exact inf_le_right.trans (le_infi₂ $ λ l hl, principal_mono.2 $ Ici_subset_Ioi.2 hl) end -lemma nhds_within_Iic_eq'' [topological_space α] [partial_order α] [order_topology α] (a : α) : +lemma nhds_within_Iic_eq'' [topological_space α] [preorder α] [order_topology α] (a : α) : 𝓝[≤] a = (⨅ l < a, 𝓟 (Ioi l)) ⊓ 𝓟 (Iic a) := nhds_within_Ici_eq'' (to_dual a) -lemma nhds_within_Ici_eq' [topological_space α] [partial_order α] [order_topology α] {a : α} +lemma nhds_within_Ici_eq' [topological_space α] [preorder α] [order_topology α] {a : α} (ha : ∃ u, a < u) : 𝓝[≥] a = ⨅ u (hu : a < u), 𝓟 (Ico a u) := by simp only [nhds_within_Ici_eq'', binfi_inf ha, inf_principal, Iio_inter_Ici] -lemma nhds_within_Iic_eq' [topological_space α] [partial_order α] [order_topology α] {a : α} +lemma nhds_within_Iic_eq' [topological_space α] [preorder α] [order_topology α] {a : α} (ha : ∃ l, l < a) : 𝓝[≤] a = ⨅ l < a, 𝓟 (Ioc l a) := by simp only [nhds_within_Iic_eq'', binfi_inf ha, inf_principal, Ioi_inter_Iic] @@ -973,11 +973,11 @@ lemma nhds_within_Iic_basis [topological_space α] [linear_order α] [order_topo [no_min_order α] (a : α) : (𝓝[≤] a).has_basis (λ l, l < a) (λ l, Ioc l a) := nhds_within_Iic_basis' (exists_lt a) -lemma nhds_top_order [topological_space α] [partial_order α] [order_top α] [order_topology α] : +lemma nhds_top_order [topological_space α] [preorder α] [order_top α] [order_topology α] : 𝓝 (⊤:α) = (⨅l (h₂ : l < ⊤), 𝓟 (Ioi l)) := by simp [nhds_eq_order (⊤:α)] -lemma nhds_bot_order [topological_space α] [partial_order α] [order_bot α] [order_topology α] : +lemma nhds_bot_order [topological_space α] [preorder α] [order_bot α] [order_topology α] : 𝓝 (⊥:α) = (⨅l (h₂ : ⊥ < l), 𝓟 (Iio l)) := by simp [nhds_eq_order (⊥:α)] @@ -1004,7 +1004,7 @@ lemma nhds_bot_basis_Iic [topological_space α] [linear_order α] [order_bot α] (𝓝 ⊥).has_basis (λ a : α, ⊥ < a) Iic := @nhds_top_basis_Ici αᵒᵈ _ _ _ _ _ _ -lemma tendsto_nhds_top_mono [topological_space β] [partial_order β] [order_top β] [order_topology β] +lemma tendsto_nhds_top_mono [topological_space β] [preorder β] [order_top β] [order_topology β] {l : filter α} {f g : α → β} (hf : tendsto f l (𝓝 ⊤)) (hg : f ≤ᶠ[l] g) : tendsto g l (𝓝 ⊤) := begin @@ -1013,17 +1013,17 @@ begin filter_upwards [hf x hx, hg] with _ using lt_of_lt_of_le, end -lemma tendsto_nhds_bot_mono [topological_space β] [partial_order β] [order_bot β] [order_topology β] +lemma tendsto_nhds_bot_mono [topological_space β] [preorder β] [order_bot β] [order_topology β] {l : filter α} {f g : α → β} (hf : tendsto f l (𝓝 ⊥)) (hg : g ≤ᶠ[l] f) : tendsto g l (𝓝 ⊥) := @tendsto_nhds_top_mono α βᵒᵈ _ _ _ _ _ _ _ hf hg -lemma tendsto_nhds_top_mono' [topological_space β] [partial_order β] [order_top β] +lemma tendsto_nhds_top_mono' [topological_space β] [preorder β] [order_top β] [order_topology β] {l : filter α} {f g : α → β} (hf : tendsto f l (𝓝 ⊤)) (hg : f ≤ g) : tendsto g l (𝓝 ⊤) := tendsto_nhds_top_mono hf (eventually_of_forall hg) -lemma tendsto_nhds_bot_mono' [topological_space β] [partial_order β] [order_bot β] +lemma tendsto_nhds_bot_mono' [topological_space β] [preorder β] [order_bot β] [order_topology β] {l : filter α} {f g : α → β} (hf : tendsto f l (𝓝 ⊥)) (hg : g ≤ f) : tendsto g l (𝓝 ⊥) := tendsto_nhds_bot_mono hf (eventually_of_forall hg) From c0dbeca1a7a7f4959cdf6b2817629bafbf1547a0 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 5 Jan 2023 16:16:13 +0000 Subject: [PATCH 0150/1291] chore(data/{int,nat}/modeq): Rationalise lemma names (#17902) Quite a few `modeq` lemmas are still called `nat.modeq.modeq_something` or `int.modeq.modeq_something`. I'm deleting the duplicated `modeq`, so that lemmas (usually) become `nat.modeq.something` and `int.modeq.something`. Also add `nat.modeq.eq_of_lt_of_lt` as a corollary to `nat.modeq.eq_of_abs_lt`. --- src/data/int/modeq.lean | 23 +++++---- src/data/nat/modeq.lean | 96 +++++++++++++++++++------------------ src/number_theory/pell.lean | 14 +++--- 3 files changed, 68 insertions(+), 65 deletions(-) diff --git a/src/data/int/modeq.lean b/src/data/int/modeq.lean index 8c2d99cf0c206..493111c330de6 100644 --- a/src/data/int/modeq.lean +++ b/src/data/int/modeq.lean @@ -66,15 +66,14 @@ begin exact exists_congr (λ t, sub_eq_iff_eq_add'), end -theorem modeq.dvd : a ≡ b [ZMOD n] → n ∣ b - a := modeq_iff_dvd.1 -theorem modeq_of_dvd : n ∣ b - a → a ≡ b [ZMOD n] := modeq_iff_dvd.2 +alias modeq_iff_dvd ↔ modeq.dvd modeq_of_dvd theorem mod_modeq (a n) : a % n ≡ a [ZMOD n] := mod_mod _ _ namespace modeq -protected theorem modeq_of_dvd (d : m ∣ n) (h : a ≡ b [ZMOD n]) : a ≡ b [ZMOD m] := -modeq_iff_dvd.2 $ d.trans h.dvd +protected lemma of_dvd (d : m ∣ n) (h : a ≡ b [ZMOD n]) : a ≡ b [ZMOD m] := +modeq_of_dvd $ d.trans h.dvd protected theorem mul_left' (hc : 0 ≤ c) (h : a ≡ b [ZMOD n]) : c * a ≡ c * b [ZMOD (c * n)] := or.cases_on hc.lt_or_eq (λ hc, @@ -124,9 +123,9 @@ h.sub modeq.rfl protected theorem mul_left (c : ℤ) (h : a ≡ b [ZMOD n]) : c * a ≡ c * b [ZMOD n] := or.cases_on (le_total 0 c) -(λ hc, (h.mul_left' hc).modeq_of_dvd (dvd_mul_left _ _) ) +(λ hc, (h.mul_left' hc).of_dvd (dvd_mul_left _ _) ) (λ hc, by rw [← neg_neg c, neg_mul, neg_mul _ b]; - exact ((h.mul_left' $ neg_nonneg.2 hc).modeq_of_dvd (dvd_mul_left _ _)).neg) + exact ((h.mul_left' $ neg_nonneg.2 hc).of_dvd (dvd_mul_left _ _)).neg) protected theorem mul_right (c : ℤ) (h : a ≡ b [ZMOD n]) : a * c ≡ b * c [ZMOD n] := by { rw [mul_comm a, mul_comm b], exact h.mul_left c } @@ -141,11 +140,11 @@ begin exact h.mul hd, end -theorem of_modeq_mul_left (m : ℤ) (h : a ≡ b [ZMOD m * n]) : a ≡ b [ZMOD n] := +theorem of_mul_left (m : ℤ) (h : a ≡ b [ZMOD m * n]) : a ≡ b [ZMOD n] := by rw [modeq_iff_dvd] at *; exact (dvd_mul_left n m).trans h -theorem of_modeq_mul_right (m : ℤ) : a ≡ b [ZMOD n * m] → a ≡ b [ZMOD n] := -mul_comm m n ▸ of_modeq_mul_left _ +theorem of_mul_right (m : ℤ) : a ≡ b [ZMOD n * m] → a ≡ b [ZMOD n] := +mul_comm m n ▸ of_mul_left _ end modeq @@ -163,7 +162,7 @@ lemma modeq_and_modeq_iff_modeq_mul {a b m n : ℤ} (hmn : m.nat_abs.coprime n.n refine hmn.mul_dvd_of_dvd_of_dvd _ _; rw [← coe_nat_dvd, nat_abs_dvd, dvd_nat_abs]; tauto end, -λ h, ⟨h.of_modeq_mul_right _, h.of_modeq_mul_left _⟩⟩ +λ h, ⟨h.of_mul_right _, h.of_mul_left _⟩⟩ lemma gcd_a_modeq (a b : ℕ) : (a : ℤ) * nat.gcd_a a b ≡ nat.gcd a b [ZMOD b] := by { rw [← add_zero ((a : ℤ) * _), nat.gcd_eq_gcd_ab], @@ -197,9 +196,9 @@ let ⟨z, hz1, hz2, hz3⟩ := exists_unique_equiv a hb in @[simp] lemma mod_mul_right_mod (a b c : ℤ) : a % (b * c) % b = a % b := -(mod_modeq _ _).of_modeq_mul_right _ +(mod_modeq _ _).of_mul_right _ @[simp] lemma mod_mul_left_mod (a b c : ℤ) : a % (b * c) % c = a % c := -(mod_modeq _ _).of_modeq_mul_left _ +(mod_modeq _ _).of_mul_left _ end int diff --git a/src/data/nat/modeq.lean b/src/data/nat/modeq.lean index d0510320fe193..e3340f3d69167 100644 --- a/src/data/nat/modeq.lean +++ b/src/data/nat/modeq.lean @@ -61,8 +61,7 @@ theorem modeq_iff_dvd : a ≡ b [MOD n] ↔ (n:ℤ) ∣ b - a := by rw [modeq, eq_comm, ← int.coe_nat_inj', int.coe_nat_mod, int.coe_nat_mod, int.mod_eq_mod_iff_mod_sub_eq_zero, int.dvd_iff_mod_eq_zero] -protected theorem modeq.dvd : a ≡ b [MOD n] → (n:ℤ) ∣ b - a := modeq_iff_dvd.1 -theorem modeq_of_dvd : (n:ℤ) ∣ b - a → a ≡ b [MOD n] := modeq_iff_dvd.2 +alias modeq_iff_dvd ↔ modeq.dvd modeq_of_dvd /-- A variant of `modeq_iff_dvd` with `nat` divisibility -/ theorem modeq_iff_dvd' (h : a ≤ b) : a ≡ b [MOD n] ↔ n ∣ b - a := @@ -72,14 +71,14 @@ theorem mod_modeq (a n) : a % n ≡ a [MOD n] := mod_mod _ _ namespace modeq -protected theorem modeq_of_dvd (d : m ∣ n) (h : a ≡ b [MOD n]) : a ≡ b [MOD m] := +protected theorem of_dvd (d : m ∣ n) (h : a ≡ b [MOD n]) : a ≡ b [MOD m] := modeq_of_dvd ((int.coe_nat_dvd.2 d).trans h.dvd) protected theorem mul_left' (c : ℕ) (h : a ≡ b [MOD n]) : c * a ≡ c * b [MOD (c * n)] := by unfold modeq at *; rw [mul_mod_mul_left, mul_mod_mul_left, h] protected theorem mul_left (c : ℕ) (h : a ≡ b [MOD n]) : c * a ≡ c * b [MOD n] := -(h.mul_left' _ ).modeq_of_dvd (dvd_mul_left _ _) +(h.mul_left' _ ).of_dvd (dvd_mul_left _ _) protected theorem mul_right' (c : ℕ) (h : a ≡ b [MOD n]) : a * c ≡ b * c [MOD (n * c)] := by rw [mul_comm a, mul_comm b, mul_comm n]; exact h.mul_left' c @@ -128,6 +127,9 @@ by { rw [add_comm a, add_comm b] at h₂, exact h₁.add_left_cancel h₂ } protected theorem add_right_cancel' (c : ℕ) (h : a + c ≡ b + c [MOD n]) : a ≡ b [MOD n] := modeq.rfl.add_right_cancel h +/-- Cancel left multiplication on both sides of the `≡` and in the modulus. + +For cancelling left multiplication in the modulus, see `nat.modeq.of_mul_left`. -/ protected theorem mul_left_cancel' {a b c m : ℕ} (hc : c ≠ 0) : c * a ≡ c * b [MOD c * m] → a ≡ b [MOD m] := by simp [modeq_iff_dvd, ←mul_sub, mul_dvd_mul_iff_left (by simp [hc] : (c : ℤ) ≠ 0)] @@ -136,6 +138,9 @@ protected theorem mul_left_cancel_iff' {a b c m : ℕ} (hc : c ≠ 0) : c * a ≡ c * b [MOD c * m] ↔ a ≡ b [MOD m] := ⟨modeq.mul_left_cancel' hc, modeq.mul_left' _⟩ +/-- Cancel right multiplication on both sides of the `≡` and in the modulus. + +For cancelling right multiplication in the modulus, see `nat.modeq.of_mul_right`. -/ protected theorem mul_right_cancel' {a b c m : ℕ} (hc : c ≠ 0) : a * c ≡ b * c [MOD m * c] → a ≡ b [MOD m] := by simp [modeq_iff_dvd, ←sub_mul, mul_dvd_mul_iff_right (by simp [hc] : (c : ℤ) ≠ 0)] @@ -144,26 +149,27 @@ protected theorem mul_right_cancel_iff' {a b c m : ℕ} (hc : c ≠ 0) : a * c ≡ b * c [MOD m * c] ↔ a ≡ b [MOD m] := ⟨modeq.mul_right_cancel' hc, modeq.mul_right' _⟩ -theorem of_modeq_mul_left (m : ℕ) (h : a ≡ b [MOD m * n]) : a ≡ b [MOD n] := +/-- Cancel left multiplication in the modulus. + +For cancelling left multiplication on both sides of the `≡`, see `nat.modeq.mul_left_cancel'`. -/ +theorem of_mul_left (m : ℕ) (h : a ≡ b [MOD m * n]) : a ≡ b [MOD n] := by { rw [modeq_iff_dvd] at *, exact (dvd_mul_left (n : ℤ) (m : ℤ)).trans h } -theorem of_modeq_mul_right (m : ℕ) : a ≡ b [MOD n * m] → a ≡ b [MOD n] := -mul_comm m n ▸ of_modeq_mul_left _ +/-- Cancel right multiplication in the modulus. + +For cancelling right multiplication on both sides of the `≡`, see `nat.modeq.mul_right_cancel'`. -/ +theorem of_mul_right (m : ℕ) : a ≡ b [MOD n * m] → a ≡ b [MOD n] := mul_comm m n ▸ of_mul_left _ end modeq -theorem modeq_one : a ≡ b [MOD 1] := modeq_of_dvd (one_dvd _) +lemma modeq_sub (h : b ≤ a) : a ≡ b [MOD a - b] := (modeq_of_dvd $ by rw [int.coe_nat_sub h]).symm -lemma modeq_sub (h : b ≤ a) : a ≡ b [MOD a - b] := -(modeq_of_dvd $ by rw [int.coe_nat_sub h]).symm +lemma modeq_one : a ≡ b [MOD 1] := modeq_of_dvd $ one_dvd _ -@[simp] lemma modeq_zero_iff {a b : ℕ} : a ≡ b [MOD 0] ↔ a = b := -by rw [nat.modeq, nat.mod_zero, nat.mod_zero] +@[simp] lemma modeq_zero_iff : a ≡ b [MOD 0] ↔ a = b := by rw [modeq, mod_zero, mod_zero] -@[simp] lemma add_modeq_left {a n : ℕ} : n + a ≡ a [MOD n] := -by rw [nat.modeq, nat.add_mod_left] -@[simp] lemma add_modeq_right {a n : ℕ} : a + n ≡ a [MOD n] := -by rw [nat.modeq, nat.add_mod_right] +@[simp] lemma add_modeq_left : n + a ≡ a [MOD n] := by rw [modeq, add_mod_left] +@[simp] lemma add_modeq_right : a + n ≡ a [MOD n] := by rw [modeq, add_mod_right] namespace modeq @@ -174,33 +180,36 @@ lemma le_of_lt_add (h1 : a ≡ b [MOD m]) (h2 : a < b + m) : a ≤ b := lemma add_le_of_lt (h1 : a ≡ b [MOD m]) (h2 : a < b) : a + m ≤ b := le_of_lt_add (add_modeq_right.trans h1) (add_lt_add_right h2 m) -lemma dvd_iff_of_modeq_of_dvd {a b d m : ℕ} (h : a ≡ b [MOD m]) (hdm : d ∣ m) : - d ∣ a ↔ d ∣ b := +lemma dvd_iff (h : a ≡ b [MOD m]) (hdm : d ∣ m) : d ∣ a ↔ d ∣ b := begin simp only [←modeq_zero_iff_dvd], - replace h := h.modeq_of_dvd hdm, + replace h := h.of_dvd hdm, exact ⟨h.symm.trans, h.trans⟩, end -lemma gcd_eq_of_modeq {a b m : ℕ} (h : a ≡ b [MOD m]) : gcd a m = gcd b m := +lemma gcd_eq (h : a ≡ b [MOD m]) : gcd a m = gcd b m := begin have h1 := gcd_dvd_right a m, have h2 := gcd_dvd_right b m, exact dvd_antisymm - (dvd_gcd ((dvd_iff_of_modeq_of_dvd h h1).mp (gcd_dvd_left a m)) h1) - (dvd_gcd ((dvd_iff_of_modeq_of_dvd h h2).mpr (gcd_dvd_left b m)) h2), + (dvd_gcd ((h.dvd_iff h1).mp (gcd_dvd_left a m)) h1) + (dvd_gcd ((h.dvd_iff h2).mpr (gcd_dvd_left b m)) h2), end -lemma eq_of_modeq_of_abs_lt {a b m : ℕ} (h : a ≡ b [MOD m]) (h2 : | (b:ℤ) - a | < m) : a = b := +lemma eq_of_abs_lt (h : a ≡ b [MOD m]) (h2 : |(b - a : ℤ)| < m) : a = b := begin apply int.coe_nat_inj, rw [eq_comm, ←sub_eq_zero], exact int.eq_zero_of_abs_lt_dvd (modeq_iff_dvd.mp h) h2, end +lemma eq_of_lt_of_lt (h : a ≡ b [MOD m]) (ha : a < m) (hb : b < m) : a = b := +h.eq_of_abs_lt $ abs_sub_lt_iff.2 + ⟨(sub_le_self _ $ int.coe_nat_nonneg _).trans_lt $ cast_lt.2 hb, + (sub_le_self _ $ int.coe_nat_nonneg _).trans_lt $ cast_lt.2 ha⟩ + /-- To cancel a common factor `c` from a `modeq` we must divide the modulus `m` by `gcd m c` -/ -lemma modeq_cancel_left_div_gcd {a b c m : ℕ} (hm : 0 < m) (h : c * a ≡ c * b [MOD m]) : - a ≡ b [MOD m / gcd m c] := +lemma cancel_left_div_gcd (hm : 0 < m) (h : c * a ≡ c * b [MOD m]) : a ≡ b [MOD m / gcd m c] := begin let d := gcd m c, have hmd := gcd_dvd_left m c, @@ -217,35 +226,30 @@ begin nat.div_self (gcd_pos_of_pos_left c hm)] }, end -lemma modeq_cancel_right_div_gcd {a b c m : ℕ} (hm : 0 < m) (h : a * c ≡ b * c [MOD m]) : - a ≡ b [MOD m / gcd m c] := -by { apply modeq_cancel_left_div_gcd hm, simpa [mul_comm] using h } +lemma cancel_right_div_gcd (hm : 0 < m) (h : a * c ≡ b * c [MOD m]) : a ≡ b [MOD m / gcd m c] := +by { apply cancel_left_div_gcd hm, simpa [mul_comm] using h } -lemma modeq_cancel_left_div_gcd' {a b c d m : ℕ} (hm : 0 < m) (hcd : c ≡ d [MOD m]) - (h : c * a ≡ d * b [MOD m]) : +lemma cancel_left_div_gcd' (hm : 0 < m) (hcd : c ≡ d [MOD m]) (h : c * a ≡ d * b [MOD m]) : a ≡ b [MOD m / gcd m c] := -modeq_cancel_left_div_gcd hm (h.trans (modeq.mul_right b hcd).symm) +(h.trans (modeq.mul_right b hcd).symm).cancel_left_div_gcd hm -lemma modeq_cancel_right_div_gcd' {a b c d m : ℕ} (hm : 0 < m) (hcd : c ≡ d [MOD m]) - (h : a * c ≡ b * d [MOD m]) : +lemma cancel_right_div_gcd' (hm : 0 < m) (hcd : c ≡ d [MOD m]) (h : a * c ≡ b * d [MOD m]) : a ≡ b [MOD m / gcd m c] := -by { apply modeq_cancel_left_div_gcd' hm hcd, simpa [mul_comm] using h } +hcd.cancel_left_div_gcd' hm $ by simpa [mul_comm] using h /-- A common factor that's coprime with the modulus can be cancelled from a `modeq` -/ -lemma modeq_cancel_left_of_coprime {a b c m : ℕ} (hmc : gcd m c = 1) (h : c * a ≡ c * b [MOD m]) : - a ≡ b [MOD m] := +lemma cancel_left_of_coprime (hmc : gcd m c = 1) (h : c * a ≡ c * b [MOD m]) : a ≡ b [MOD m] := begin rcases m.eq_zero_or_pos with rfl | hm, { simp only [gcd_zero_left] at hmc, simp only [gcd_zero_left, hmc, one_mul, modeq_zero_iff] at h, subst h }, - simpa [hmc] using modeq_cancel_left_div_gcd hm h + simpa [hmc] using h.cancel_left_div_gcd hm end /-- A common factor that's coprime with the modulus can be cancelled from a `modeq` -/ -lemma modeq_cancel_right_of_coprime {a b c m : ℕ} (hmc : gcd m c = 1) (h : a * c ≡ b * c [MOD m]) : - a ≡ b [MOD m] := -by { apply modeq_cancel_left_of_coprime hmc, simpa [mul_comm] using h } +lemma cancel_right_of_coprime (hmc : gcd m c = 1) (h : a * c ≡ b * c [MOD m]) : a ≡ b [MOD m] := +cancel_left_of_coprime hmc $ by simpa [mul_comm] using h end modeq @@ -306,22 +310,22 @@ lemma modeq_and_modeq_iff_modeq_mul {a b m n : ℕ} (hmn : coprime m n) : rw [nat.modeq_iff_dvd, ← int.dvd_nat_abs, int.coe_nat_dvd], exact hmn.mul_dvd_of_dvd_of_dvd h.1 h.2 end, -λ h, ⟨h.of_modeq_mul_right _, h.of_modeq_mul_left _⟩⟩ +λ h, ⟨h.of_mul_right _, h.of_mul_left _⟩⟩ lemma coprime_of_mul_modeq_one (b : ℕ) {a n : ℕ} (h : a * b ≡ 1 [MOD n]) : coprime a n := begin obtain ⟨g, hh⟩ := nat.gcd_dvd_right a n, rw [nat.coprime_iff_gcd_eq_one, ← nat.dvd_one, ← nat.modeq_zero_iff_dvd], - calc 1 ≡ a * b [MOD a.gcd n] : nat.modeq.of_modeq_mul_right g (hh.subst h).symm + calc 1 ≡ a * b [MOD a.gcd n] : nat.modeq.of_mul_right g (hh.subst h).symm ... ≡ 0 * b [MOD a.gcd n] : (nat.modeq_zero_iff_dvd.mpr (nat.gcd_dvd_left _ _)).mul_right b ... = 0 : by rw zero_mul, end @[simp] lemma mod_mul_right_mod (a b c : ℕ) : a % (b * c) % b = a % b := -(mod_modeq _ _).of_modeq_mul_right _ +(mod_modeq _ _).of_mul_right _ @[simp] lemma mod_mul_left_mod (a b c : ℕ) : a % (b * c) % c = a % c := -(mod_modeq _ _).of_modeq_mul_left _ +(mod_modeq _ _).of_mul_left _ lemma div_mod_eq_mod_mul_div (a b c : ℕ) : a / b % c = a % (b * c) / b := if hb0 : b = 0 then by simp [hb0] @@ -414,11 +418,11 @@ by rw [mul_add, two_mul_odd_div_two hm1, mul_left_comm, two_mul_odd_div_two hn1, tsub_add_cancel_of_le (le_mul_of_one_le_right (nat.zero_le _) hn0)] lemma odd_of_mod_four_eq_one {n : ℕ} : n % 4 = 1 → n % 2 = 1 := -by simpa [modeq, show 2 * 2 = 4, by norm_num] using @modeq.of_modeq_mul_left 2 n 1 2 +by simpa [modeq, show 2 * 2 = 4, by norm_num] using @modeq.of_mul_left 2 n 1 2 lemma odd_of_mod_four_eq_three {n : ℕ} : n % 4 = 3 → n % 2 = 1 := by simpa [modeq, show 2 * 2 = 4, by norm_num, show 3 % 4 = 3, by norm_num] - using @modeq.of_modeq_mul_left 2 n 3 2 + using @modeq.of_mul_left 2 n 3 2 /-- A natural number is odd iff it has residue `1` or `3` mod `4`-/ lemma odd_mod_four_iff {n : ℕ} : n % 2 = 1 ↔ n % 4 = 1 ∨ n % 4 = 3 := diff --git a/src/number_theory/pell.lean b/src/number_theory/pell.lean index ffdd65bd0783a..2ce2c8408b372 100644 --- a/src/number_theory/pell.lean +++ b/src/number_theory/pell.lean @@ -313,7 +313,7 @@ theorem xy_modeq_yn (n) : (hx.mul_right _ ).add $ modeq_zero_iff_dvd.2 $ by rw pow_succ'; exact mul_dvd_mul_right (dvd_mul_of_dvd_right (modeq_zero_iff_dvd.1 $ - (hy.modeq_of_dvd $ by simp [pow_succ']).trans $ modeq_zero_iff_dvd.2 $ + (hy.of_dvd $ by simp [pow_succ']).trans $ modeq_zero_iff_dvd.2 $ by simp [-mul_comm, -mul_assoc]) _) _, have R : xn (n * k) * yn n + yn (n * k) * xn n ≡ xn n^k * yn n + k * xn n^k * yn n [MOD yn n^3], from @@ -327,7 +327,7 @@ theorem xy_modeq_yn (n) : theorem ysq_dvd_yy (n) : yn n * yn n ∣ yn (n * yn n) := modeq_zero_iff_dvd.1 $ - ((xy_modeq_yn n (yn n)).right.modeq_of_dvd $ by simp [pow_succ]).trans + ((xy_modeq_yn n (yn n)).right.of_dvd $ by simp [pow_succ]).trans (modeq_zero_iff_dvd.2 $ by simp [mul_dvd_mul_left, mul_assoc]) theorem dvd_of_ysq_dvd {n t} (h : yn n * yn n ∣ yn t) : yn n ∣ t := @@ -337,7 +337,7 @@ let ⟨k, ke⟩ := nt in have yn n ∣ k * (xn n)^(k-1), from nat.dvd_of_mul_dvd_mul_right (strict_mono_y n0l) $ modeq_zero_iff_dvd.1 $ by have xm := (xy_modeq_yn a1 n k).right; rw ← ke at xm; exact - (xm.modeq_of_dvd $ by simp [pow_succ]).symm.trans h.modeq_zero_nat, + (xm.of_dvd $ by simp [pow_succ]).symm.trans h.modeq_zero_nat, by rw ke; exact dvd_mul_of_dvd_right (((xy_coprime _ _).pow_left _).symm.dvd_of_dvd_mul_right this) _ @@ -641,7 +641,7 @@ theorem matiyasevic {a k x y} : (∃ a1 : 1 < a, xn a1 k = x ∧ yn a1 k = y) have 4 * y ∣ b - 1, from int.coe_nat_dvd.1 $ by rw int.coe_nat_sub (le_of_lt b1); exact bm1.symm.dvd, - (yn_modeq_a_sub_one _ _).modeq_of_dvd this, + (yn_modeq_a_sub_one _ _).of_dvd this, ⟨ky, or.inr ⟨u, v, s, t, b, pell_eq _ _, pell_eq _ _, pell_eq _ _, b1, bm1, ba, vp, yv, sx, tk⟩⟩, λ⟨a1, ky, o⟩, ⟨a1, match o with @@ -666,7 +666,7 @@ theorem matiyasevic {a k x y} : (∃ a1 : 1 < a, xn a1 k = x ∧ yn a1 k = y) have jk : j ≡ k [MOD 4 * yn a1 i], from have 4 * yn a1 i ∣ b - 1, from int.coe_nat_dvd.1 $ by rw int.coe_nat_sub (le_of_lt b1); exact bm1.symm.dvd, - ((yn_modeq_a_sub_one b1 _).modeq_of_dvd this).symm.trans tk, + ((yn_modeq_a_sub_one b1 _).of_dvd this).symm.trans tk, have ki : k + i < 4 * yn a1 i, from lt_of_le_of_lt (add_le_add ky (yn_ge_n a1 i)) $ by rw ← two_mul; exact nat.mul_lt_mul_of_pos_right dec_trivial (strict_mono_y a1 ipos), @@ -675,9 +675,9 @@ theorem matiyasevic {a k x y} : (∃ a1 : 1 < a, xn a1 k = x ∧ yn a1 k = y) (modeq_of_xn_modeq a1 ipos iln this).resolve_right $ λ (ji : j + i ≡ 0 [MOD 4 * n]), not_le_of_gt ki $ nat.le_of_dvd (lt_of_lt_of_le ipos $ nat.le_add_left _ _) $ modeq_zero_iff_dvd.1 $ (jk.symm.add_right i).trans $ - ji.modeq_of_dvd yd, + ji.of_dvd yd, by have : i % (4 * yn a1 i) = k % (4 * yn a1 i) := - (ji.modeq_of_dvd yd).symm.trans jk; + (ji.of_dvd yd).symm.trans jk; rwa [nat.mod_eq_of_lt (lt_of_le_of_lt (nat.le_add_left _ _) ki), nat.mod_eq_of_lt (lt_of_le_of_lt (nat.le_add_right _ _) ki)] at this end From 26f081a2fb920140ed5bc5cc5344e84bcc7cb2b2 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 5 Jan 2023 18:39:26 +0000 Subject: [PATCH 0151/1291] feat(data/list/infix): `list.slice` is a sublist of the original (#18067) --- src/data/list/infix.lean | 11 +++++++++++ 1 file changed, 11 insertions(+) diff --git a/src/data/list/infix.lean b/src/data/list/infix.lean index 1d8d58e50fd64..7f94d09d1f190 100644 --- a/src/data/list/infix.lean +++ b/src/data/list/infix.lean @@ -212,6 +212,17 @@ lemma drop_subset (n) (l : list α) : drop n l ⊆ l := (drop_sublist n l).subse lemma mem_of_mem_take (h : a ∈ l.take n) : a ∈ l := take_subset n l h lemma mem_of_mem_drop (h : a ∈ l.drop n) : a ∈ l := drop_subset n l h +lemma slice_sublist (n m : ℕ) (l : list α) : l.slice n m <+ l := +begin + rw list.slice_eq, + conv_rhs {rw ←list.take_append_drop n l}, + rw [list.append_sublist_append_left, add_comm, list.drop_add], + exact list.drop_sublist _ _, +end +lemma slice_subset (n m : ℕ) (l : list α) : l.slice n m ⊆ l := (slice_sublist n m l).subset +lemma mem_of_mem_slice {n m : ℕ} {l : list α} {a : α} (h : a ∈ l.slice n m) : a ∈ l := +slice_subset n m l h + lemma take_while_prefix (p : α → Prop) [decidable_pred p] : l.take_while p <+: l := ⟨l.drop_while p, take_while_append_drop p l⟩ From 00f4ab49e7d5139216e0b3daad15fffa504897ab Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Fri, 6 Jan 2023 01:37:19 +0000 Subject: [PATCH 0152/1291] chore(*): add mathlib4 synchronization comments (#18072) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.tropical.basic` * `algebra.tropical.lattice` * `data.list.basic` * `data.list.infix` * `data.list.lemmas` * `data.list.min_max` * `data.list.palindrome` * `data.list.tfae` * `data.set.intervals.ord_connected_component` * `order.concept` --- src/algebra/tropical/basic.lean | 3 +++ src/algebra/tropical/lattice.lean | 3 +++ src/data/list/basic.lean | 3 +++ src/data/list/infix.lean | 3 +++ src/data/list/lemmas.lean | 3 +++ src/data/list/min_max.lean | 3 +++ src/data/list/palindrome.lean | 3 +++ src/data/list/tfae.lean | 3 +++ src/data/set/intervals/ord_connected_component.lean | 3 +++ src/order/concept.lean | 3 +++ 10 files changed, 30 insertions(+) diff --git a/src/algebra/tropical/basic.lean b/src/algebra/tropical/basic.lean index 28206249952f0..463ad1f35fec0 100644 --- a/src/algebra/tropical/basic.lean +++ b/src/algebra/tropical/basic.lean @@ -12,6 +12,9 @@ import algebra.order.monoid.min_max # Tropical algebraic structures +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines algebraic structures of the (min-)tropical numbers, up to the tropical semiring. Some basic lemmas about conversion from the base type `R` to `tropical R` are provided, as well as the expected implementations of tropical addition and tropical multiplication. diff --git a/src/algebra/tropical/lattice.lean b/src/algebra/tropical/lattice.lean index 83b1d1606b71d..544bba5582e4a 100644 --- a/src/algebra/tropical/lattice.lean +++ b/src/algebra/tropical/lattice.lean @@ -10,6 +10,9 @@ import order.conditionally_complete_lattice.basic # Order on tropical algebraic structure +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the orders induced on tropical algebraic structures by the underlying type. ## Main declarations diff --git a/src/data/list/basic.lean b/src/data/list/basic.lean index a1213cece0149..051098553ce50 100644 --- a/src/data/list/basic.lean +++ b/src/data/list/basic.lean @@ -7,6 +7,9 @@ import data.nat.order.basic /-! # Basic properties of lists + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ open function nat (hiding one_pos) diff --git a/src/data/list/infix.lean b/src/data/list/infix.lean index 7f94d09d1f190..89bb5ba243d14 100644 --- a/src/data/list/infix.lean +++ b/src/data/list/infix.lean @@ -8,6 +8,9 @@ import data.list.basic /-! # Prefixes, subfixes, infixes +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves properties about * `list.prefix`: `l₁` is a prefix of `l₂` if `l₂` starts with `l₁`. * `list.subfix`: `l₁` is a subfix of `l₂` if `l₂` ends with `l₁`. diff --git a/src/data/list/lemmas.lean b/src/data/list/lemmas.lean index 144a707cf5c44..afc0d0070e9eb 100644 --- a/src/data/list/lemmas.lean +++ b/src/data/list/lemmas.lean @@ -8,6 +8,9 @@ import data.list.basic /-! # Some lemmas about lists involving sets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Split out from `data.list.basic` to reduce its dependencies. -/ diff --git a/src/data/list/min_max.lean b/src/data/list/min_max.lean index ec342ac73ae67..734c8f1320542 100644 --- a/src/data/list/min_max.lean +++ b/src/data/list/min_max.lean @@ -8,6 +8,9 @@ import data.list.basic /-! # Minimum and maximum of lists +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions The main definitions are `argmax`, `argmin`, `minimum` and `maximum` for lists. diff --git a/src/data/list/palindrome.lean b/src/data/list/palindrome.lean index e38cc03ddc2d8..74d22c24122dd 100644 --- a/src/data/list/palindrome.lean +++ b/src/data/list/palindrome.lean @@ -9,6 +9,9 @@ import data.list.basic /-! # Palindromes +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This module defines *palindromes*, lists which are equal to their reverse. The main result is the `palindrome` inductive type, and its associated `palindrome.rec_on` induction diff --git a/src/data/list/tfae.lean b/src/data/list/tfae.lean index c6838bb452d29..28cd9ad514de8 100644 --- a/src/data/list/tfae.lean +++ b/src/data/list/tfae.lean @@ -8,6 +8,9 @@ import data.list.basic /-! # The Following Are Equivalent +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file allows to state that all propositions in a list are equivalent. It is used by `tactic.tfae`. `tfae l` means `∀ x ∈ l, ∀ y ∈ l, x ↔ y`. This is equivalent to `pairwise (↔) l`. diff --git a/src/data/set/intervals/ord_connected_component.lean b/src/data/set/intervals/ord_connected_component.lean index 3c9878d4ef42f..3bc77087f89db 100644 --- a/src/data/set/intervals/ord_connected_component.lean +++ b/src/data/set/intervals/ord_connected_component.lean @@ -9,6 +9,9 @@ import tactic.wlog /-! # Order connected components of a set +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define `set.ord_connected_component s x` to be the set of `y` such that `set.interval x y ⊆ s` and prove some basic facts about this definition. At the moment of writing, this construction is used only to prove that any linear order with order topology is a T₅ space, diff --git a/src/order/concept.lean b/src/order/concept.lean index 0327b0d58dc01..feea7695528d4 100644 --- a/src/order/concept.lean +++ b/src/order/concept.lean @@ -8,6 +8,9 @@ import data.set.lattice /-! # Formal concept analysis +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines concept lattices. A concept of a relation `r : α → β → Prop` is a pair of sets `s : set α` and `t : set β` such that `s` is the set of all `a : α` that are related to all elements of `t`, and `t` is the set of all `b : β` that are related to all elements of `s`. From 18a5306c091183ac90884daa9373fa3b178e8607 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Martin=20Dvo=C5=99=C3=A1k?= Date: Fri, 6 Jan 2023 05:33:22 +0000 Subject: [PATCH 0153/1291] docs(data/nat/choose/sum): binomial theorem is for commuting elements (#18071) Co-authored-by: madvorak --- src/data/nat/choose/sum.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/data/nat/choose/sum.lean b/src/data/nat/choose/sum.lean index 23a502f0b1918..fb9d4b7488df9 100644 --- a/src/data/nat/choose/sum.lean +++ b/src/data/nat/choose/sum.lean @@ -31,7 +31,7 @@ variables [semiring R] {x y : R} (h : commute x y) (n : ℕ) include h -/-- A version of the **binomial theorem** for noncommutative semirings. -/ +/-- A version of the **binomial theorem** for commuting elements in noncommutative semirings. -/ theorem add_pow : (x + y) ^ n = ∑ m in range (n + 1), x ^ m * y ^ (n - m) * choose n m := begin From 7a89b1aed52bcacbcc4a8ad515e72c5c07268940 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 6 Jan 2023 07:59:14 +0000 Subject: [PATCH 0154/1291] feat(data/int/cast/lemmas): `a.nat_mod b < b` (#17896) The modulus of an integer by a natural is strictly less than the natural. Co-authored-by: Eric Wieser --- src/data/int/cast/lemmas.lean | 6 ++++++ 1 file changed, 6 insertions(+) diff --git a/src/data/int/cast/lemmas.lean b/src/data/int/cast/lemmas.lean index 6598a6553085a..75356b0146c7e 100644 --- a/src/data/int/cast/lemmas.lean +++ b/src/data/int/cast/lemmas.lean @@ -36,6 +36,12 @@ def of_nat_hom : ℕ →+* ℤ := ⟨coe, rfl, int.of_nat_mul, rfl, int.of_nat_a lemma coe_nat_succ_pos (n : ℕ) : 0 < (n.succ : ℤ) := int.coe_nat_pos.2 (succ_pos n) +lemma to_nat_lt {a : ℤ} {b : ℕ} (hb : b ≠ 0) : a.to_nat < b ↔ a < b := +by { rw [←to_nat_lt_to_nat, to_nat_coe_nat], exact coe_nat_pos.2 hb.bot_lt } + +lemma nat_mod_lt {a : ℤ} {b : ℕ} (hb : b ≠ 0) : a.nat_mod b < b := +(to_nat_lt hb).2 $ mod_lt_of_pos _ $ coe_nat_pos.2 hb.bot_lt + section cast @[simp, norm_cast] theorem cast_mul [non_assoc_ring α] : ∀ m n, ((m * n : ℤ) : α) = m * n := From 2ed7e4aec72395b6a7c3ac4ac7873a7a43ead17c Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9mi=20Bottinelli?= Date: Fri, 6 Jan 2023 10:58:12 +0000 Subject: [PATCH 0155/1291] feat(analysis/bounded_variation): lower semicontinuity of `variation_on` (#18058) --- src/analysis/bounded_variation.lean | 51 +++++++++++++++++++++++++++-- 1 file changed, 48 insertions(+), 3 deletions(-) diff --git a/src/analysis/bounded_variation.lean b/src/analysis/bounded_variation.lean index f2368959ba082..9adb29ee3322c 100644 --- a/src/analysis/bounded_variation.lean +++ b/src/analysis/bounded_variation.lean @@ -41,9 +41,8 @@ it possible to use the complete linear order structure of `ℝ≥0∞`. The proo more tedious with an `ℝ`-valued or `ℝ≥0`-valued variation, since one would always need to check that the sets one uses are nonempty and bounded above as these are only conditionally complete. -/ - -open_locale big_operators nnreal ennreal -open set measure_theory +open_locale big_operators nnreal ennreal topological_space uniform_convergence +open set measure_theory filter variables {α β : Type*} [linear_order α] [linear_order β] {E F : Type*} [pseudo_emetric_space E] [pseudo_emetric_space F] @@ -207,6 +206,52 @@ begin simp [u, edist_comm], end +lemma lower_continuous_aux {ι : Type*} {F : ι → α → E} {p : filter ι} + {f : α → E} {s : set α} (Ffs : ∀ x ∈ s, tendsto (λ i, F i x) p (𝓝 (f x))) + {v : ℝ≥0∞} (hv : v < evariation_on f s) : ∀ᶠ (n : ι) in p, v < evariation_on (F n) s := +begin + obtain ⟨⟨n, ⟨u, um, us⟩⟩, hlt⟩ : + ∃ (p : ℕ × {u : ℕ → α // monotone u ∧ ∀ i, u i ∈ s}), + v < ∑ i in finset.range p.1, edist (f ((p.2 : ℕ → α) (i+1))) (f ((p.2 : ℕ → α) i)) := + lt_supr_iff.mp hv, + have : tendsto (λ j, ∑ (i : ℕ) in finset.range n, edist (F j (u (i + 1))) (F j (u i))) + p (𝓝 (∑ (i : ℕ) in finset.range n, edist (f (u (i + 1))) (f (u i)))), + { apply tendsto_finset_sum, + exact λ i hi, tendsto.edist (Ffs (u i.succ) (us i.succ)) (Ffs (u i) (us i)) }, + exact (eventually_gt_of_tendsto_gt hlt this).mono + (λ i h, lt_of_lt_of_le h (sum_le (F i) n um us)), +end + +/-- +The map `λ f, evariation_on f s` is lower semicontinuous for pointwise convergence *on `s`*. +Pointwise convergence on `s` is encoded here as uniform convergence on the family consisting of the +singletons of elements of `s`. +-/ +@[protected] +lemma lower_semicontinuous (s : set α) : + lower_semicontinuous (λ f : α →ᵤ[s.image singleton] E, evariation_on f s) := +begin + intro f, + apply @lower_continuous_aux _ _ _ _ (uniform_on_fun α E (s.image singleton)) id (𝓝 f) f s _, + simpa only [uniform_on_fun.tendsto_iff_tendsto_uniformly_on, mem_image, forall_exists_index, + and_imp, forall_apply_eq_imp_iff₂, + tendsto_uniformly_on_singleton_iff_tendsto] using @tendsto_id _ (𝓝 f), +end + +/-- +The map `λ f, evariation_on f s` is lower semicontinuous for uniform convergence on `s`. +-/ +lemma lower_semicontinuous_uniform_on (s : set α) : + lower_semicontinuous (λ f : α →ᵤ[{s}] E, evariation_on f s) := +begin + intro f, + apply @lower_continuous_aux _ _ _ _ (uniform_on_fun α E {s}) id (𝓝 f) f s _, + have := @tendsto_id _ (𝓝 f), + rw uniform_on_fun.tendsto_iff_tendsto_uniformly_on at this, + simp_rw ←tendsto_uniformly_on_singleton_iff_tendsto, + exact λ x xs, ((this s rfl).mono (singleton_subset_iff.mpr xs)), +end + lemma _root_.has_bounded_variation_on.dist_le {E : Type*} [pseudo_metric_space E] {f : α → E} {s : set α} (h : has_bounded_variation_on f s) {x y : α} (hx : x ∈ s) (hy : y ∈ s) : dist (f x) (f y) ≤ (evariation_on f s).to_real := From 730513c79ef70a35c2fe70d2f64855d23b52352f Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 6 Jan 2023 13:16:04 +0000 Subject: [PATCH 0156/1291] feat(group_theory/perm/basic): `mul_left` is a monoid hom (#17900) `equiv.mul_left` is a monoid homomorphism. --- src/algebra/hom/iterate.lean | 7 +-- src/group_theory/perm/basic.lean | 75 ++++++++++++++++++++++++++++++-- 2 files changed, 72 insertions(+), 10 deletions(-) diff --git a/src/algebra/hom/iterate.lean b/src/algebra/hom/iterate.lean index b9e900512f21a..4bd68c2cdae3c 100644 --- a/src/algebra/hom/iterate.lean +++ b/src/algebra/hom/iterate.lean @@ -3,9 +3,7 @@ Copyright (c) 2020 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ - -import logic.function.iterate -import group_theory.perm.basic +import algebra.group_power.lemmas import group_theory.group_action.opposite /-! @@ -144,9 +142,6 @@ f.to_add_monoid_hom.iterate_map_zsmul n m x end ring_hom -lemma equiv.perm.coe_pow {α : Type*} (f : equiv.perm α) (n : ℕ) : ⇑(f ^ n) = (f^[n]) := -hom_coe_pow _ rfl (λ _ _, rfl) _ _ - --what should be the namespace for this section? section monoid diff --git a/src/group_theory/perm/basic.lean b/src/group_theory/perm/basic.lean index 844bbea80a7ad..7ae9350c243e3 100644 --- a/src/group_theory/perm/basic.lean +++ b/src/group_theory/perm/basic.lean @@ -4,9 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura, Mario Carneiro -/ import algebra.group.pi -import algebra.group_power.lemmas import algebra.group.prod -import logic.function.iterate +import algebra.hom.iterate /-! # The group of permutations (self-equivalences) of a type `α` @@ -66,9 +65,11 @@ lemma mul_def (f g : perm α) : f * g = g.trans f := rfl lemma inv_def (f : perm α) : f⁻¹ = f.symm := rfl -@[simp] lemma coe_mul (f g : perm α) : ⇑(f * g) = f ∘ g := rfl +@[simp, norm_cast] lemma coe_one : ⇑(1 : perm α) = id := rfl +@[simp, norm_cast] lemma coe_mul (f g : perm α) : ⇑(f * g) = f ∘ g := rfl -@[simp] lemma coe_one : ⇑(1 : perm α) = id := rfl +@[norm_cast] lemma coe_pow (f : perm α) (n : ℕ) : ⇑(f ^ n) = (f^[n]) := +hom_coe_pow _ rfl (λ _ _, rfl) _ _ lemma eq_inv_iff_eq {f : perm α} {x y : α} : x = f⁻¹ y ↔ f x = y := f.eq_symm_apply @@ -426,4 +427,70 @@ equiv.ext $ λ n, by { simp only [swap_apply_def, perm.mul_apply], split_ifs; cc end swap +section add_group +variables [add_group α] (a b : α) + +@[simp] lemma add_left_zero : equiv.add_left (0 : α) = 1 := ext zero_add +@[simp] lemma add_right_zero : equiv.add_right (0 : α) = 1 := ext add_zero + +@[simp] lemma add_left_add : equiv.add_left (a + b) = equiv.add_left a * equiv.add_left b := +ext $ add_assoc _ _ + +@[simp] lemma add_right_add : equiv.add_right (a + b) = equiv.add_right b * equiv.add_right a := +ext $ λ _, (add_assoc _ _ _).symm + +@[simp] lemma inv_add_left : (equiv.add_left a)⁻¹ = equiv.add_left (-a) := equiv.coe_inj.1 rfl +@[simp] lemma inv_add_right : (equiv.add_right a)⁻¹ = equiv.add_right (-a) := equiv.coe_inj.1 rfl + +@[simp] lemma pow_add_left (n : ℕ) : equiv.add_left a ^ n = equiv.add_left (n • a) := +by { ext, simp [perm.coe_pow] } + +@[simp] lemma pow_add_right (n : ℕ) : equiv.add_right a ^ n = equiv.add_right (n • a) := +by { ext, simp [perm.coe_pow] } + +@[simp] lemma zpow_add_left (n : ℤ) : equiv.add_left a ^ n = equiv.add_left (n • a) := +(map_zsmul (⟨equiv.add_left, add_left_zero, add_left_add⟩ : α →+ additive (perm α)) _ _).symm + +@[simp] lemma zpow_add_right (n : ℤ) : equiv.add_right a ^ n = equiv.add_right (n • a) := +@zpow_add_left αᵃᵒᵖ _ _ _ + +end add_group + +section group +variables [group α] (a b : α) + +@[simp, to_additive] lemma mul_left_one : equiv.mul_left (1 : α) = 1 := ext one_mul +@[simp, to_additive] lemma mul_right_one : equiv.mul_right (1 : α) = 1 := ext mul_one + +@[simp, to_additive] +lemma mul_left_mul : equiv.mul_left (a * b) = equiv.mul_left a * equiv.mul_left b := +ext $ mul_assoc _ _ + +@[simp, to_additive] +lemma mul_right_mul : equiv.mul_right (a * b) = equiv.mul_right b * equiv.mul_right a := +ext $ λ _, (mul_assoc _ _ _).symm + +@[simp, to_additive inv_add_left] +lemma inv_mul_left : (equiv.mul_left a)⁻¹ = equiv.mul_left a⁻¹ := equiv.coe_inj.1 rfl +@[simp, to_additive inv_add_right] +lemma inv_mul_right : (equiv.mul_right a)⁻¹ = equiv.mul_right a⁻¹ := equiv.coe_inj.1 rfl + +@[simp, to_additive pow_add_left] +lemma pow_mul_left (n : ℕ) : equiv.mul_left a ^ n = equiv.mul_left (a ^ n) := +by { ext, simp [perm.coe_pow] } + +@[simp, to_additive pow_add_right] +lemma pow_mul_right (n : ℕ) : equiv.mul_right a ^ n = equiv.mul_right (a ^ n) := +by { ext, simp [perm.coe_pow] } + +@[simp, to_additive zpow_add_left] +lemma zpow_mul_left (n : ℤ) : equiv.mul_left a ^ n = equiv.mul_left (a ^ n) := +(map_zpow (⟨equiv.mul_left, mul_left_one, mul_left_mul⟩ : α →* perm α) _ _).symm + +@[simp, to_additive zpow_add_right] +lemma zpow_mul_right : ∀ n : ℤ, equiv.mul_right a ^ n = equiv.mul_right (a ^ n) +| (int.of_nat n) := by simp +| (int.neg_succ_of_nat n) := by simp + +end group end equiv From 792a2a264169d64986541c6f8f7e3bbb6acb6295 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 6 Jan 2023 13:16:05 +0000 Subject: [PATCH 0157/1291] feat(logic/function/iterate): Cancelling iterates (#17956) We can cancel iterates of an injective function. --- src/logic/function/iterate.lean | 15 ++++++++++++++- 1 file changed, 14 insertions(+), 1 deletion(-) diff --git a/src/logic/function/iterate.lean b/src/logic/function/iterate.lean index 1f383d7e9bde4..7ee0de3cba916 100644 --- a/src/logic/function/iterate.lean +++ b/src/logic/function/iterate.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import logic.function.conjugate +import tactic.alias /-! # Iterations of a function @@ -151,7 +152,7 @@ lemma iterate.rec_zero (p : α → Sort*) {f : α → α} (h : ∀ a, p a → p iterate.rec p h ha 0 = ha := rfl -variable {f} +variables {f} {m n : ℕ} {a : α} theorem left_inverse.iterate {g : α → α} (hg : left_inverse g f) (n : ℕ) : left_inverse (g^[n]) (f^[n]) := @@ -167,6 +168,18 @@ lemma iterate_comm (f : α → α) (m n : ℕ) : f^[n]^[m] = (f^[m]^[n]) := lemma iterate_commute (m n : ℕ) : commute (λ f : α → α, f^[m]) (λ f, f^[n]) := λ f, iterate_comm f m n +lemma iterate_add_eq_iterate (hf : injective f) : f^[m + n] a = (f^[n] a) ↔ (f^[m] a) = a := +iff.trans (by rw [←iterate_add_apply, nat.add_comm]) (hf.iterate n).eq_iff + +alias iterate_add_eq_iterate ↔ iterate_cancel_of_add _ + +lemma iterate_cancel (hf : injective f) (ha : f^[m] a = (f^[n] a)) : f^[m - n] a = a := +begin + cases le_total m n, + { simp [nat.sub_eq_zero_of_le h] }, + { exact iterate_cancel_of_add hf (by rwa nat.sub_add_cancel h) } +end + end function namespace list From d6aad9528ddcac270ed35c6f7b5f1d8af25341d6 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 6 Jan 2023 16:29:35 +0000 Subject: [PATCH 0158/1291] =?UTF-8?q?feat(order/lattice):=20`a=20=E2=8A=94?= =?UTF-8?q?=20b=20=3D=20a=20=E2=8A=93=20b=20=E2=86=94=20a=20=3D=20b`=20(#1?= =?UTF-8?q?7966)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit and `a ⊓ b = c ∧ a ⊔ b = c ↔ a = c ∧ b = c` --- src/order/lattice.lean | 23 +++++++++++++---------- 1 file changed, 13 insertions(+), 10 deletions(-) diff --git a/src/order/lattice.lean b/src/order/lattice.lean index e41a99100a96b..df7d86f946fb9 100644 --- a/src/order/lattice.lean +++ b/src/order/lattice.lean @@ -504,20 +504,23 @@ variables [lattice α] {a b c d : α} lemma inf_le_sup : a ⊓ b ≤ a ⊔ b := inf_le_left.trans le_sup_left -@[simp] lemma inf_lt_sup : a ⊓ b < a ⊔ b ↔ a ≠ b := -begin - split, - { rintro H rfl, simpa using H }, - { refine λ Hne, lt_iff_le_and_ne.2 ⟨inf_le_sup, λ Heq, Hne _⟩, - refine le_antisymm _ _, - exacts [le_sup_left.trans (Heq.symm.trans_le inf_le_right), - le_sup_right.trans (Heq.symm.trans_le inf_le_left)] } -end - @[simp] lemma sup_le_inf : a ⊔ b ≤ a ⊓ b ↔ a = b := ⟨λ h, le_antisymm (le_sup_left.trans $ h.trans inf_le_right) (le_sup_right.trans $ h.trans inf_le_left), by { rintro rfl, simp }⟩ +@[simp] lemma inf_eq_sup : a ⊓ b = a ⊔ b ↔ a = b := by rw [←inf_le_sup.ge_iff_eq, sup_le_inf] +@[simp] lemma sup_eq_inf : a ⊔ b = a ⊓ b ↔ a = b := eq_comm.trans inf_eq_sup +@[simp] lemma inf_lt_sup : a ⊓ b < a ⊔ b ↔ a ≠ b := by rw [inf_le_sup.lt_iff_ne, ne.def, inf_eq_sup] + +lemma inf_eq_and_sup_eq_iff : a ⊓ b = c ∧ a ⊔ b = c ↔ a = c ∧ b = c := +begin + refine ⟨λ h, _, _⟩, + { obtain rfl := sup_eq_inf.1 (h.2.trans h.1.symm), + simpa using h }, + { rintro ⟨rfl, rfl⟩, + exact ⟨inf_idem, sup_idem⟩ } +end + /-! #### Distributivity laws -/ From 55d224c38461be1e8e4363247dd110137c24a4ff Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Fri, 6 Jan 2023 16:29:37 +0000 Subject: [PATCH 0159/1291] feat(ring_theory/polynomial/cyclotomic/eval): golf, add lemmas (#18022) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit * Add `nat.le_self_pow`, `polynomial.sub_one_pow_totient_le_cyclotomic_eval`, `polynomial.cyclotomic_eval_le_add_one_pow_totient`, `polynomial.sub_one_pow_totient_lt_nat_abs_cyclotomic_eval`, `zmod.order_of_units_dvd_card_sub_one`, and `zmod.order_of_dvd_card_sub_one`. * Rename `polynomial.cyclotomic_eval_lt_sub_one_pow_totient` to `polynomial.cyclotomic_eval_lt_add_one_pow_totient`. * Replace `nat.exists_prime_ge_modeq_one` with `nat.exists_prime_gt_modeq_one`. * Assume `≠ 0` instead of `0 <` in `nat.exists_prime_ge_modeq_one` etc. * Golf some proofs. Mathlib 4 version: leanprover-community/mathlib4#1273 --- archive/imo/imo2008_q3.lean | 2 +- src/data/nat/factorization/basic.lean | 3 +- src/data/nat/log.lean | 2 +- src/data/nat/pow.lean | 4 ++ src/field_theory/finite/basic.lean | 8 +++ src/number_theory/primes_congruent_one.lean | 65 +++++++++---------- .../polynomial/cyclotomic/eval.lean | 55 ++++++++-------- 7 files changed, 73 insertions(+), 66 deletions(-) diff --git a/archive/imo/imo2008_q3.lean b/archive/imo/imo2008_q3.lean index 07addcdb172bb..3f8d05c99f40f 100644 --- a/archive/imo/imo2008_q3.lean +++ b/archive/imo/imo2008_q3.lean @@ -81,7 +81,7 @@ theorem imo2008_q3 : ∀ N : ℕ, ∃ n : ℕ, n ≥ N ∧ ∃ p : ℕ, nat.prime p ∧ p ∣ n ^ 2 + 1 ∧ (p : ℝ) > 2 * n + sqrt(2 * n) := begin intro N, - obtain ⟨p, hpp, hineq₁, hpmod4⟩ := nat.exists_prime_ge_modeq_one (N ^ 2 + 21) zero_lt_four, + obtain ⟨p, hpp, hineq₁, hpmod4⟩ := nat.exists_prime_gt_modeq_one (N ^ 2 + 20) four_ne_zero, obtain ⟨n, hnat, hreal⟩ := p_lemma p hpp hpmod4 (by linarith [hineq₁, nat.zero_le (N ^ 2)]), have hineq₂ : n ^ 2 + 1 ≥ p := nat.le_of_dvd (n ^ 2).succ_pos hnat, diff --git a/src/data/nat/factorization/basic.lean b/src/data/nat/factorization/basic.lean index 9fa837a0571fe..950bdb874bac2 100644 --- a/src/data/nat/factorization/basic.lean +++ b/src/data/nat/factorization/basic.lean @@ -743,8 +743,7 @@ begin by_cases ha1 : a = 1, { rw [ha1, mul_one], exact hp p n hp' hn }, - refine h (p^n) a ((hp'.one_lt).trans_le (le_self_pow (prime.one_lt hp').le hn.ne')) - _ _ (hp _ _ hp' hn) hPa, + refine h (p^n) a ((hp'.one_lt).trans_le (le_self_pow hn.ne' _)) _ _ (hp _ _ hp' hn) hPa, { contrapose! hpa, simp [lt_one_iff.1 (lt_of_le_of_ne hpa ha1)] }, simpa [hn, prime.coprime_iff_not_dvd hp'], diff --git a/src/data/nat/log.lean b/src/data/nat/log.lean index a2e46d0a572de..97c2d88aa2d8f 100644 --- a/src/data/nat/log.lean +++ b/src/data/nat/log.lean @@ -70,7 +70,7 @@ begin { have b_pos : 0 < b := zero_le_one.trans_lt hb, rw [succ_eq_add_one, add_le_add_iff_right, ←ih (y / b) (div_lt_self hy.bot_lt hb) (nat.div_pos h.1 b_pos).ne', le_div_iff_mul_le b_pos, pow_succ'] }, - { exact iff_of_false (λ hby, h ⟨(le_self_pow hb.le x.succ_ne_zero).trans hby, hb⟩) + { exact iff_of_false (λ hby, h ⟨(le_self_pow x.succ_ne_zero _).trans hby, hb⟩) (not_succ_le_zero _) } end diff --git a/src/data/nat/pow.lean b/src/data/nat/pow.lean index 219b088bccaf2..4b2ac6720c4e6 100644 --- a/src/data/nat/pow.lean +++ b/src/data/nat/pow.lean @@ -35,6 +35,10 @@ pow_lt_pow H h lemma pow_lt_pow_succ {p : ℕ} (h : 1 < p) (n : ℕ) : p^n < p^(n+1) := pow_lt_pow_of_lt_right h n.lt_succ_self +lemma le_self_pow {n : ℕ} (hn : n ≠ 0) : ∀ m : ℕ, m ≤ m ^ n +| 0 := zero_le _ +| (m + 1) := _root_.le_self_pow dec_trivial hn + lemma lt_pow_self {p : ℕ} (h : 1 < p) : ∀ n : ℕ, n < p ^ n | 0 := by simp [zero_lt_one] | (n+1) := calc diff --git a/src/field_theory/finite/basic.lean b/src/field_theory/finite/basic.lean index 66eda6e0c97c2..d75cbf1233591 100644 --- a/src/field_theory/finite/basic.lean +++ b/src/field_theory/finite/basic.lean @@ -417,6 +417,14 @@ theorem pow_card_sub_one_eq_one {p : ℕ} [fact p.prime] {a : zmod p} (ha : a a ^ (p - 1) = 1 := by { have h := pow_card_sub_one_eq_one a ha, rwa zmod.card p at h } +theorem order_of_units_dvd_card_sub_one {p : ℕ} [fact p.prime] (u : (zmod p)ˣ) : + order_of u ∣ p - 1 := +order_of_dvd_of_pow_eq_one $ units_pow_card_sub_one_eq_one _ _ + +theorem order_of_dvd_card_sub_one {p : ℕ} [fact p.prime] {a : zmod p} (ha : a ≠ 0) : + order_of a ∣ p - 1 := +order_of_dvd_of_pow_eq_one $ pow_card_sub_one_eq_one ha + open polynomial lemma expand_card {p : ℕ} [fact p.prime] (f : polynomial (zmod p)) : diff --git a/src/number_theory/primes_congruent_one.lean b/src/number_theory/primes_congruent_one.lean index ced195fc722d0..98cb1db00ec5d 100644 --- a/src/number_theory/primes_congruent_one.lean +++ b/src/number_theory/primes_congruent_one.lean @@ -17,60 +17,55 @@ We prove that, for any positive `k : ℕ`, there are infinitely many primes `p` namespace nat open polynomial nat filter +open_locale nat -/-- For any positive `k : ℕ` there are infinitely many primes `p` such that `p ≡ 1 [MOD k]`. -/ -lemma exists_prime_ge_modeq_one {k : ℕ} (n : ℕ) (hpos : 0 < k) : - ∃ (p : ℕ), nat.prime p ∧ n ≤ p ∧ p ≡ 1 [MOD k] := +/-- For any positive `k : ℕ` there exists an arbitrarily large prime `p` such that +`p ≡ 1 [MOD k]`. -/ +lemma exists_prime_gt_modeq_one {k : ℕ} (n : ℕ) (hk0 : k ≠ 0) : + ∃ (p : ℕ), nat.prime p ∧ n < p ∧ p ≡ 1 [MOD k] := begin - let b := 3 * (k * n.factorial), + rcases (one_le_iff_ne_zero.2 hk0).eq_or_lt with rfl | hk1, + { rcases exists_infinite_primes (n + 1) with ⟨p, hnp, hp⟩, + exact ⟨p, hp, hnp, modeq_one⟩ }, + let b := k * n!, have hgt : 1 < (eval ↑b (cyclotomic k ℤ)).nat_abs, - { have hkey : ∀ l : ℕ, 2 < 3 * (l.succ * n.factorial) := λ l, lt_mul_of_lt_of_one_le - (2 : ℕ).lt_succ_self (le_mul_of_le_of_one_le (nat.succ_pos _) n.factorial_pos), - rcases k with _ | _ | k, - { simpa using hpos, }, - { simp only [one_mul, int.coe_nat_mul, int.coe_nat_succ, int.coe_nat_zero, zero_add, - cyclotomic_one, eval_sub, eval_X, eval_one], - convert int.nat_abs_lt_nat_abs_of_nonneg_of_lt int.one_nonneg _, - rw lt_sub_iff_add_lt, - specialize hkey 0, - norm_cast, - rwa one_mul at hkey, }, - calc 1 ≤ _ : by { rw le_tsub_iff_left (one_le_two.trans (hkey _).le), exact (hkey _).le, } - ... < _ : sub_one_lt_nat_abs_cyclotomic_eval (one_lt_succ_succ k) - (one_lt_two.trans (hkey k.succ)).ne.symm, }, + { rcases le_iff_exists_add'.1 hk1.le with ⟨k, rfl⟩, + have hb : 2 ≤ b := le_mul_of_le_of_one_le hk1 n.factorial_pos, + calc 1 ≤ b - 1 : le_tsub_of_add_le_left hb + ... < (eval (b : ℤ) (cyclotomic (k + 1) ℤ)).nat_abs : + sub_one_lt_nat_abs_cyclotomic_eval hk1 (succ_le_iff.1 hb).ne' }, let p := min_fac (eval ↑b (cyclotomic k ℤ)).nat_abs, haveI hprime : fact p.prime := ⟨min_fac_prime (ne_of_lt hgt).symm⟩, have hroot : is_root (cyclotomic k (zmod p)) (cast_ring_hom (zmod p) b), { rw [is_root.def, ← map_cyclotomic_int k (zmod p), eval_map, coe_cast_ring_hom, - ← int.cast_coe_nat, ← int.coe_cast_ring_hom, eval₂_hom, int.coe_cast_ring_hom, + ← int.cast_coe_nat, ← int.coe_cast_ring_hom, eval₂_hom, int.coe_cast_ring_hom, zmod.int_coe_zmod_eq_zero_iff_dvd _ _], apply int.dvd_nat_abs.1, exact_mod_cast min_fac_dvd (eval ↑b (cyclotomic k ℤ)).nat_abs }, - refine ⟨p, hprime.1, _, _⟩, - { by_contra habs, - exact (prime.dvd_iff_not_coprime hprime.1).1 - (dvd_factorial (min_fac_pos _) (le_of_not_ge habs)) - (coprime_of_root_cyclotomic hpos hroot).symm.coprime_mul_left_right.coprime_mul_left_right }, - { have hdiv := order_of_dvd_of_pow_eq_one (zmod.units_pow_card_sub_one_eq_one p - (zmod.unit_of_coprime b (coprime_of_root_cyclotomic hpos hroot))), - have : ¬p ∣ k := hprime.1.coprime_iff_not_dvd.1 - (coprime_of_root_cyclotomic hpos hroot).symm.coprime_mul_left_right.coprime_mul_right_right, - haveI := ne_zero.of_not_dvd (zmod p) this, + have hpb : ¬(p ∣ b) := + hprime.1.coprime_iff_not_dvd.1 (coprime_of_root_cyclotomic hk0.bot_lt hroot).symm, + refine ⟨p, hprime.1, not_le.1 $ λ habs, _, _⟩, + { exact hpb (dvd_mul_of_dvd_right (dvd_factorial (min_fac_pos _) habs) _) }, + { have hdiv : order_of (b : zmod p) ∣ p - 1 := + zmod.order_of_dvd_card_sub_one (mt (char_p.cast_eq_zero_iff _ _ _).1 hpb), + haveI : ne_zero (k : zmod p) := + ne_zero.of_not_dvd (zmod p) (λ hpk, hpb (dvd_mul_of_dvd_left hpk _)), have : k = order_of (b : zmod p) := (is_root_cyclotomic_iff.mp hroot).eq_order_of, - rw [←order_of_units, zmod.coe_unit_of_coprime, ←this] at hdiv, + rw [← this] at hdiv, exact ((modeq_iff_dvd' hprime.1.pos).2 hdiv).symm } end -lemma frequently_at_top_modeq_one {k : ℕ} (hpos : 0 < k) : +lemma frequently_at_top_modeq_one {k : ℕ} (hk0 : k ≠ 0) : ∃ᶠ p in at_top, nat.prime p ∧ p ≡ 1 [MOD k] := begin refine frequently_at_top.2 (λ n, _), - obtain ⟨p, hp⟩ := exists_prime_ge_modeq_one n hpos, - exact ⟨p, ⟨hp.2.1, hp.1, hp.2.2⟩⟩ + obtain ⟨p, hp⟩ := exists_prime_gt_modeq_one n hk0, + exact ⟨p, ⟨hp.2.1.le, hp.1, hp.2.2⟩⟩ end -lemma infinite_set_of_prime_modeq_one {k : ℕ} (hpos : 0 < k) : +/-- For any positive `k : ℕ` there are infinitely many primes `p` such that `p ≡ 1 [MOD k]`. -/ +lemma infinite_set_of_prime_modeq_one {k : ℕ} (hk0 : k ≠ 0) : set.infinite {p : ℕ | nat.prime p ∧ p ≡ 1 [MOD k]} := -frequently_at_top_iff_infinite.1 (frequently_at_top_modeq_one hpos) +frequently_at_top_iff_infinite.1 (frequently_at_top_modeq_one hk0) end nat diff --git a/src/ring_theory/polynomial/cyclotomic/eval.lean b/src/ring_theory/polynomial/cyclotomic/eval.lean index 1da812833eafd..3a5bec45b320a 100644 --- a/src/ring_theory/polynomial/cyclotomic/eval.lean +++ b/src/ring_theory/polynomial/cyclotomic/eval.lean @@ -232,7 +232,13 @@ begin simpa only [hq'.le, real.coe_to_nnreal', max_eq_left, sub_nonneg] using hex }, end -lemma cyclotomic_eval_lt_sub_one_pow_totient {n : ℕ} {q : ℝ} (hn' : 3 ≤ n) (hq' : 1 < q) : +lemma sub_one_pow_totient_le_cyclotomic_eval {q : ℝ} (hq' : 1 < q) : + ∀ n, (q - 1) ^ totient n ≤ (cyclotomic n ℝ).eval q +| 0 := by simp only [totient_zero, pow_zero, cyclotomic_zero, eval_one] +| 1 := by simp only [totient_one, pow_one, cyclotomic_one, eval_sub, eval_X, eval_one] +| (n + 2) := (sub_one_pow_totient_lt_cyclotomic_eval dec_trivial hq').le + +lemma cyclotomic_eval_lt_add_one_pow_totient {n : ℕ} {q : ℝ} (hn' : 3 ≤ n) (hq' : 1 < q) : (cyclotomic n ℝ).eval q < (q + 1) ^ totient n := begin have hn : 0 < n := pos_of_gt hn', @@ -297,33 +303,28 @@ begin exact ⟨ζ, hζ, by simp [hhζ]⟩ }, end -lemma sub_one_lt_nat_abs_cyclotomic_eval {n : ℕ} {q : ℕ} (hn' : 1 < n) (hq' : q ≠ 1) : - q - 1 < ((cyclotomic n ℤ).eval ↑q).nat_abs := +lemma cyclotomic_eval_le_add_one_pow_totient {q : ℝ} (hq' : 1 < q) : + ∀ n, (cyclotomic n ℝ).eval q ≤ (q + 1) ^ totient n +| 0 := by simp +| 1 := by simp [add_assoc, add_nonneg, zero_le_one] +| 2 := by simp +| (n + 3) := (cyclotomic_eval_lt_add_one_pow_totient dec_trivial hq').le + +lemma sub_one_pow_totient_lt_nat_abs_cyclotomic_eval {n : ℕ} {q : ℕ} (hn' : 1 < n) (hq : q ≠ 1) : + (q - 1) ^ totient n < ((cyclotomic n ℤ).eval ↑q).nat_abs := begin - rcases q with _ | _ | q, - iterate 2 - { rw [pos_iff_ne_zero, ne.def, int.nat_abs_eq_zero], - intro h, - have := degree_eq_one_of_irreducible_of_root (cyclotomic.irreducible (pos_of_gt hn')) h, - rw [degree_cyclotomic, with_top.coe_eq_one, totient_eq_one_iff] at this, - rcases this with rfl|rfl; simpa using h }, - suffices : (q.succ : ℝ) < (eval (↑q + 1 + 1) (cyclotomic n ℤ)).nat_abs, - { exact_mod_cast this }, - calc _ ≤ ((q + 2 - 1) ^ n.totient : ℝ) : _ - ... < _ : _, - { norm_num, - convert pow_mono (by simp : 1 ≤ (q : ℝ) + 1) (totient_pos (pos_of_gt hn') : 1 ≤ n.totient), - { simp }, - { ring }, }, - convert sub_one_pow_totient_lt_cyclotomic_eval (show 2 ≤ n, by linarith) - (show (1 : ℝ) < q + 2, by {norm_cast, linarith}), - norm_cast, - erw cyclotomic.eval_apply (q + 2 : ℤ) n (algebra_map ℤ ℝ), - simp only [int.coe_nat_succ, eq_int_cast], - norm_cast, - rw [int.coe_nat_abs_eq_normalize, int.normalize_of_nonneg], - simp only [int.coe_nat_succ], - exact cyclotomic_nonneg n (by linarith), + rcases hq.lt_or_lt.imp_left nat.lt_one_iff.mp with rfl | hq', + { rw [zero_tsub, zero_pow (nat.totient_pos (pos_of_gt hn')), pos_iff_ne_zero, int.nat_abs_ne_zero, + nat.cast_zero, ← coeff_zero_eq_eval_zero, cyclotomic_coeff_zero _ hn'], + exact one_ne_zero }, + rw [← @nat.cast_lt ℝ, nat.cast_pow, nat.cast_sub hq'.le, nat.cast_one, int.cast_nat_abs], + refine (sub_one_pow_totient_lt_cyclotomic_eval hn' (nat.one_lt_cast.2 hq')).trans_le _, + exact (cyclotomic.eval_apply (q : ℤ) n (algebra_map ℤ ℝ)).trans_le (le_abs_self _) end +lemma sub_one_lt_nat_abs_cyclotomic_eval {n : ℕ} {q : ℕ} (hn' : 1 < n) (hq : q ≠ 1) : + q - 1 < ((cyclotomic n ℤ).eval ↑q).nat_abs := +calc q - 1 ≤ (q - 1) ^ totient n : nat.le_self_pow (nat.totient_pos $ pos_of_gt hn').ne' _ +... < ((cyclotomic n ℤ).eval ↑q).nat_abs : sub_one_pow_totient_lt_nat_abs_cyclotomic_eval hn' hq + end polynomial From c85b4d15cfdac3abffce2f449e228a0cf0326912 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Fri, 6 Jan 2023 16:29:38 +0000 Subject: [PATCH 0160/1291] chore(algebra/continued_fractions): golf 2 proofs (#18070) --- .../computation/approximations.lean | 57 ++++++------------- 1 file changed, 18 insertions(+), 39 deletions(-) diff --git a/src/algebra/continued_fractions/computation/approximations.lean b/src/algebra/continued_fractions/computation/approximations.lean index 15194e3e7af1f..99a04f2f26d94 100644 --- a/src/algebra/continued_fractions/computation/approximations.lean +++ b/src/algebra/continued_fractions/computation/approximations.lean @@ -275,19 +275,10 @@ lemma le_of_succ_succ_nth_continuants_aux_b {b : K} (nth_part_denom_eq : (of v).partial_denominators.nth n = some b) : b * ((of v).continuants_aux $ n + 1).b ≤ ((of v).continuants_aux $ n + 2).b := begin - set g := of v with g_eq, - obtain ⟨gp_n, nth_s_eq, gpnb_eq_b⟩ : ∃ gp_n, g.s.nth n = some gp_n ∧ gp_n.b = b, from - exists_s_b_of_part_denom nth_part_denom_eq, - subst gpnb_eq_b, - let conts := g.continuants_aux (n + 2), - set pconts := g.continuants_aux (n + 1) with pconts_eq, - set ppconts := g.continuants_aux n with ppconts_eq, - have h1 : 0 ≤ ppconts.b, from zero_le_of_continuants_aux_b, - have h2 : gp_n.b * pconts.b ≤ ppconts.b + gp_n.b * pconts.b, - { solve_by_elim [le_add_of_nonneg_of_le, le_refl] }, - -- use the recurrence of continuants_aux and the fact that gp_n.a = 1 - simp [h1, h2, of_part_num_eq_one (part_num_eq_s_a nth_s_eq), - generalized_continued_fraction.continuants_aux_recurrence nth_s_eq ppconts_eq pconts_eq], + obtain ⟨gp_n, nth_s_eq, rfl⟩ : ∃ gp_n, (of v).s.nth n = some gp_n ∧ gp_n.b = b, + from exists_s_b_of_part_denom nth_part_denom_eq, + simp [of_part_num_eq_one (part_num_eq_s_a nth_s_eq), zero_le_of_continuants_aux_b, + generalized_continued_fraction.continuants_aux_recurrence nth_s_eq rfl rfl] end /-- Shows that `bₙ * Bₙ ≤ Bₙ₊₁`, where `bₙ` is the `n`th partial denominator and `Bₙ₊₁` and `Bₙ` are @@ -573,33 +564,21 @@ Shows that `|v - Aₙ / Bₙ| ≤ 1 / (bₙ * Bₙ * Bₙ)`. This bound is worse -/ lemma abs_sub_convergents_le' {b : K} (nth_part_denom_eq : (of v).partial_denominators.nth n = some b) : - |v - (of v).convergents n| - ≤ 1 / (b * ((of v).denominators n) * ((of v).denominators n)) := + |v - (of v).convergents n| ≤ 1 / (b * ((of v).denominators n) * ((of v).denominators n)) := begin - let g := of v, - let B := g.denominators n, - let nB := g.denominators (n + 1), - have not_terminated_at_n : ¬g.terminated_at n, by - { have : g.partial_denominators.nth n ≠ none, by simp [nth_part_denom_eq], - exact (not_iff_not_of_iff terminated_at_iff_part_denom_none).elim_right this }, - suffices : 1 / (B * nB) ≤ (1 : K) / (b * B * B), by - { have : |v - g.convergents n| ≤ 1 / (B * nB), from abs_sub_convergents_le not_terminated_at_n, - transitivity; - assumption }, - -- derive some inequalities needed to show the claim - have zero_lt_B : 0 < B, by - { have : (fib (n + 1) : K) ≤ B, from - succ_nth_fib_le_of_nth_denom (or.inr $ - mt (terminated_stable n.pred_le) not_terminated_at_n), - exact (lt_of_lt_of_le - (by exact_mod_cast (fib_pos (lt_of_le_of_ne n.succ.zero_le n.succ_ne_zero.symm))) this) }, - have denoms_ineq : b * B * B ≤ B * nB, by - { have : b * B ≤ nB, from le_of_succ_nth_denom nth_part_denom_eq, - rwa [(mul_comm B nB), (mul_le_mul_right zero_lt_B)] }, - have : (0 : K) < b * B * B, by - { have : 0 < b, from lt_of_lt_of_le zero_lt_one (of_one_le_nth_part_denom nth_part_denom_eq), - any_goals { repeat { apply mul_pos } }; assumption }, - exact (div_le_div_of_le_left zero_le_one this denoms_ineq) + have not_terminated_at_n : ¬(of v).terminated_at n, + by simp [terminated_at_iff_part_denom_none, nth_part_denom_eq], + refine (abs_sub_convergents_le not_terminated_at_n).trans _, + -- One can show that `0 < (generalized_continued_fraction.of v).denominators n` but it's easier + -- to consider the case `(generalized_continued_fraction.of v).denominators n = 0`. + rcases zero_le_of_denom.eq_or_gt + with (hB : (generalized_continued_fraction.of v).denominators n = 0) | hB, + { simp only [hB, mul_zero, zero_mul, div_zero] }, + { apply one_div_le_one_div_of_le, + { have : 0 < b := zero_lt_one.trans_le (of_one_le_nth_part_denom nth_part_denom_eq), + apply_rules [mul_pos] }, + { conv_rhs { rw [mul_comm] }, + exact mul_le_mul_of_nonneg_right (le_of_succ_nth_denom nth_part_denom_eq) hB.le } } end end error_term From 6afc9b06856ad973f6a2619e3e8a0a8d537a58f2 Mon Sep 17 00:00:00 2001 From: Eric Rodriguez <37984851+ericrbg@users.noreply.github.com> Date: Fri, 6 Jan 2023 16:29:39 +0000 Subject: [PATCH 0161/1291] chore(analysis/normed_space/lp_equiv): fix docs (#18073) Note that [in the current docs](https://leanprover-community.github.io/mathlib_docs/analysis/normed_space/lp_equiv.html), $L^p$ shows on its own line. This is because the $L^p$ are written using double $s, which is wrong. --- src/analysis/normed_space/lp_equiv.lean | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/src/analysis/normed_space/lp_equiv.lean b/src/analysis/normed_space/lp_equiv.lean index ce69795dbb044..b705a4bd3ce8c 100644 --- a/src/analysis/normed_space/lp_equiv.lean +++ b/src/analysis/normed_space/lp_equiv.lean @@ -8,16 +8,16 @@ import analysis.normed_space.pi_Lp import topology.continuous_function.bounded /-! -# Equivalences among $$L^p$$ spaces +# Equivalences among $L^p$ spaces -In this file we collect a variety of equivalences among various $$L^p$$ spaces. In particular, +In this file we collect a variety of equivalences among various $L^p$ spaces. In particular, when `α` is a `fintype`, given `E : α → Type u` and `p : ℝ≥0∞`, there is a natural linear isometric equivalence `lp_pi_Lpₗᵢ : lp E p ≃ₗᵢ pi_Lp p E`. In addition, when `α` is a discrete topological space, the bounded continuous functions `α →ᵇ β` correspond exactly to `lp (λ _, β) ∞`. Here there can be more structure, including ring and algebra structures, and we implement these equivalences accordingly as well. -We keep this as a separate file so that the various $$L^p$$ space files don't import the others. +We keep this as a separate file so that the various $L^p$ space files don't import the others. Recall that `pi_Lp` is just a type synonym for `Π i, E i` but given a different metric and norm structure, although the topological, uniform and bornological structures coincide definitionally. From e644458573f279c9db84dc6ffc5e9fa2730e9014 Mon Sep 17 00:00:00 2001 From: Floris van Doorn Date: Fri, 6 Jan 2023 19:49:12 +0000 Subject: [PATCH 0162/1291] feat(geometry/manifold/cont_mdiff): generalize some lemmas to arbitrary charts (#17668) * This PR generalizes some smoothness results from `ext_chart_at` to arbitrary members of the maximal atlas * Useful for smooth vector bundle refactor * Motivated by the sphere eversion project --- src/geometry/manifold/cont_mdiff.lean | 202 +++++++++++------- .../manifold/local_invariant_properties.lean | 24 ++- .../smooth_manifold_with_corners.lean | 25 +++ 3 files changed, 161 insertions(+), 90 deletions(-) diff --git a/src/geometry/manifold/cont_mdiff.lean b/src/geometry/manifold/cont_mdiff.lean index 212ebe27b846d..e96446cc2f90e 100644 --- a/src/geometry/manifold/cont_mdiff.lean +++ b/src/geometry/manifold/cont_mdiff.lean @@ -71,6 +71,7 @@ variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] -- F'' is a normed space {F'' : Type*} [normed_add_comm_group F''] [normed_space 𝕜 F''] -- declare functions, sets, points and smoothness indices +{e : local_homeomorph M H} {e' : local_homeomorph M' H'} {f f₁ : M → M'} {s s₁ t : set M} {x : M} {m n : ℕ∞} /-- Property in the model space of a model with corners of being `C^n` within at set at a point, @@ -147,13 +148,13 @@ lemma cont_diff_within_at_local_invariant_prop (n : ℕ∞) : { assume y hy, simp only with mfld_simps at hy, simpa only [hy] with mfld_simps using hs hy.1 } end } -lemma cont_diff_within_at_prop_mono (n : ℕ∞) - ⦃s x t⦄ ⦃f : H → H'⦄ (hts : t ⊆ s) (h : cont_diff_within_at_prop I I' n f s x) : +lemma cont_diff_within_at_prop_mono_of_mem (n : ℕ∞) + ⦃s x t⦄ ⦃f : H → H'⦄ (hts : s ∈ 𝓝[t] x) (h : cont_diff_within_at_prop I I' n f s x) : cont_diff_within_at_prop I I' n f t x := begin - apply h.mono (λ y hy, _), - simp only with mfld_simps at hy, - simp only [hy, hts _] with mfld_simps + refine h.mono_of_mem _, + refine inter_mem _ (mem_of_superset self_mem_nhds_within $ inter_subset_right _ _), + rwa [← filter.mem_map, ← I.image_eq, I.symm_map_nhds_within_image] end lemma cont_diff_within_at_prop_id (x : H) : @@ -334,17 +335,39 @@ cont_mdiff_at_iff_target include Is I's +lemma cont_mdiff_within_at_iff_of_mem_maximal_atlas + {x : M} (he : e ∈ maximal_atlas I M) (he' : e' ∈ maximal_atlas I' M') + (hx : x ∈ e.source) (hy : f x ∈ e'.source) : + cont_mdiff_within_at I I' n f s x ↔ continuous_within_at f s x ∧ + cont_diff_within_at 𝕜 n (e'.extend I' ∘ f ∘ (e.extend I).symm) + ((e.extend I).symm ⁻¹' s ∩ range I) + (e.extend I x) := +(cont_diff_within_at_local_invariant_prop I I' n).lift_prop_within_at_indep_chart he hx he' hy + +/-- An alternative formulation of `cont_mdiff_within_at_iff_of_mem_maximal_atlas` + if the set if `s` lies in `e.source`. -/ +lemma cont_mdiff_within_at_iff_image {x : M} (he : e ∈ maximal_atlas I M) + (he' : e' ∈ maximal_atlas I' M') (hs : s ⊆ e.source) (hx : x ∈ e.source) (hy : f x ∈ e'.source) : + cont_mdiff_within_at I I' n f s x ↔ continuous_within_at f s x ∧ + cont_diff_within_at 𝕜 n (e'.extend I' ∘ f ∘ (e.extend I).symm) (e.extend I '' s) (e.extend I x) := +begin + rw [cont_mdiff_within_at_iff_of_mem_maximal_atlas he he' hx hy, and.congr_right_iff], + refine λ hf, cont_diff_within_at_congr_nhds _, + simp_rw [nhds_within_eq_iff_eventually_eq, + e.extend_symm_preimage_inter_range_eventually_eq I hs hx] +end + /-- One can reformulate smoothness within a set at a point as continuity within this set at this point, and smoothness in any chart containing that point. -/ -lemma cont_mdiff_within_at_iff_of_mem_source {x' : M} {y : M'} (hx : x' ∈ (chart_at H x).source) +lemma cont_mdiff_within_at_iff_of_mem_source + {x' : M} {y : M'} (hx : x' ∈ (chart_at H x).source) (hy : f x' ∈ (chart_at H' y).source) : cont_mdiff_within_at I I' n f s x' ↔ continuous_within_at f s x' ∧ cont_diff_within_at 𝕜 n (ext_chart_at I' y ∘ f ∘ (ext_chart_at I x).symm) ((ext_chart_at I x).symm ⁻¹' s ∩ range I) (ext_chart_at I x x') := -(cont_diff_within_at_local_invariant_prop I I' n).lift_prop_within_at_indep_chart - (structure_groupoid.chart_mem_maximal_atlas _ x) hx - (structure_groupoid.chart_mem_maximal_atlas _ y) hy +cont_mdiff_within_at_iff_of_mem_maximal_atlas + (chart_mem_maximal_atlas _ x) (chart_mem_maximal_atlas _ y) hx hy lemma cont_mdiff_within_at_iff_of_mem_source' {x' : M} {y : M'} (hx : x' ∈ (chart_at H x).source) (hy : f x' ∈ (chart_at H' y).source) : @@ -402,44 +425,29 @@ begin end omit I's -variable (I) -lemma model_with_corners.symm_continuous_within_at_comp_right_iff {X} [topological_space X] - {f : H → X} {s : set H} {x : H} : - continuous_within_at (f ∘ I.symm) (I.symm ⁻¹' s ∩ range I) (I x) ↔ continuous_within_at f s x := +include Is + +lemma cont_mdiff_within_at_iff_source_of_mem_maximal_atlas + (he : e ∈ maximal_atlas I M) (hx : x ∈ e.source) : + cont_mdiff_within_at I I' n f s x ↔ + cont_mdiff_within_at 𝓘(𝕜, E) I' n (f ∘ (e.extend I).symm) + ((e.extend I).symm ⁻¹' s ∩ range I) (e.extend I x) := begin - refine ⟨λ h, _, λ h, _⟩, - { have := h.comp I.continuous_within_at (maps_to_preimage _ _), - simp_rw [preimage_inter, preimage_preimage, I.left_inv, preimage_id', preimage_range, - inter_univ] at this, - rwa [function.comp.assoc, I.symm_comp_self] at this }, - { rw [← I.left_inv x] at h, exact h.comp I.continuous_within_at_symm (inter_subset_left _ _) } + have h2x := hx, rw [← e.extend_source I] at h2x, + simp_rw [cont_mdiff_within_at, + (cont_diff_within_at_local_invariant_prop I I' n).lift_prop_within_at_indep_chart_source + he hx, structure_groupoid.lift_prop_within_at_self_source, + e.extend_symm_continuous_within_at_comp_right_iff, cont_diff_within_at_prop_self_source, + cont_diff_within_at_prop, function.comp, e.left_inv hx, (e.extend I).left_inv h2x], + refl, end -variable {I} - -lemma ext_chart_at_symm_continuous_within_at_comp_right_iff {X} [topological_space X] {f : M → X} - {s : set M} {x x' : M} : - continuous_within_at (f ∘ (ext_chart_at I x).symm) ((ext_chart_at I x).symm ⁻¹' s ∩ range I) - (ext_chart_at I x x') ↔ - continuous_within_at (f ∘ (chart_at H x).symm) ((chart_at H x).symm ⁻¹' s) (chart_at H x x') := -by convert I.symm_continuous_within_at_comp_right_iff; refl - -include Is lemma cont_mdiff_within_at_iff_source_of_mem_source {x' : M} (hx' : x' ∈ (chart_at H x).source) : cont_mdiff_within_at I I' n f s x' ↔ cont_mdiff_within_at 𝓘(𝕜, E) I' n (f ∘ (ext_chart_at I x).symm) ((ext_chart_at I x).symm ⁻¹' s ∩ range I) (ext_chart_at I x x') := -begin - have h2x' := hx', rw [← ext_chart_at_source I] at h2x', - simp_rw [cont_mdiff_within_at, - (cont_diff_within_at_local_invariant_prop I I' n).lift_prop_within_at_indep_chart_source - (chart_mem_maximal_atlas I x) hx', structure_groupoid.lift_prop_within_at_self_source, - ext_chart_at_symm_continuous_within_at_comp_right_iff, cont_diff_within_at_prop_self_source, - cont_diff_within_at_prop, function.comp, (chart_at H x).left_inv hx', - (ext_chart_at I x).left_inv h2x'], - refl, -end +cont_mdiff_within_at_iff_source_of_mem_maximal_atlas (chart_mem_maximal_atlas I x) hx' lemma cont_mdiff_at_iff_source_of_mem_source {x' : M} (hx' : x' ∈ (chart_at H x).source) : @@ -448,23 +456,20 @@ lemma cont_mdiff_at_iff_source_of_mem_source by simp_rw [cont_mdiff_at, cont_mdiff_within_at_iff_source_of_mem_source hx', preimage_univ, univ_inter] -lemma cont_mdiff_at_ext_chart_at' {x' : M} (h : x' ∈ (chart_at H x).source) : - cont_mdiff_at I 𝓘(𝕜, E) n (ext_chart_at I x) x' := +include I's + +lemma cont_mdiff_on_iff_of_mem_maximal_atlas + (he : e ∈ maximal_atlas I M) (he' : e' ∈ maximal_atlas I' M') + (hs : s ⊆ e.source) + (h2s : maps_to f s e'.source) : + cont_mdiff_on I I' n f s ↔ continuous_on f s ∧ + cont_diff_on 𝕜 n (e'.extend I' ∘ f ∘ (e.extend I).symm) + (e.extend I '' s) := begin - refine (cont_mdiff_at_iff_of_mem_source h (mem_chart_source _ _)).mpr _, - rw [← ext_chart_at_source I] at h, - refine ⟨continuous_at_ext_chart_at' _ _ h, _⟩, - refine cont_diff_within_at_id.congr_of_eventually_eq _ _, - { refine eventually_eq_of_mem (ext_chart_at_target_mem_nhds_within' I x h) (λ x₂ hx₂, _), - simp_rw [function.comp_apply, (ext_chart_at I x).right_inv hx₂], refl }, - simp_rw [function.comp_apply, (ext_chart_at I x).right_inv ((ext_chart_at I x).maps_to h)], refl + simp_rw [continuous_on, cont_diff_on, set.ball_image_iff, ← forall_and_distrib, cont_mdiff_on], + exact forall₂_congr (λ x hx, cont_mdiff_within_at_iff_image he he' hs (hs hx) (h2s hx)) end -lemma cont_mdiff_at_ext_chart_at : cont_mdiff_at I 𝓘(𝕜, E) n (ext_chart_at I x) x := -cont_mdiff_at_ext_chart_at' $ mem_chart_source H x - -include I's - /-- If the set where you want `f` to be smooth lies entirely in a single chart, and `f` maps it into a single chart, the smoothness of `f` on that set can be expressed by purely looking in these charts. @@ -476,24 +481,8 @@ lemma cont_mdiff_on_iff_of_subset_source {x : M} {y : M'} cont_mdiff_on I I' n f s ↔ continuous_on f s ∧ cont_diff_on 𝕜 n (ext_chart_at I' y ∘ f ∘ (ext_chart_at I x).symm) (ext_chart_at I x '' s) := -begin - split, - { refine λ H, ⟨λ x hx, (H x hx).1, _⟩, - rintro _ ⟨x', hx', rfl⟩, - exact ((cont_mdiff_within_at_iff_of_mem_source (hs hx') - (h2s.image_subset $ mem_image_of_mem f hx')).mp (H _ hx')).2.mono - (maps_to_ext_chart_at I x hs).image_subset }, - { rintro ⟨h1, h2⟩ x' hx', - refine (cont_mdiff_within_at_iff_of_mem_source (hs hx') - (h2s.image_subset $ mem_image_of_mem f hx')).mpr - ⟨h1.continuous_within_at hx', _⟩, - refine (h2 _ $ mem_image_of_mem _ hx').mono_of_mem _, - rw [← ext_chart_at_source I] at hs, - rw [(ext_chart_at I x).image_eq_target_inter_inv_preimage hs], - refine inter_mem _ (ext_chart_at_preimage_mem_nhds_within' I x (hs hx') self_mem_nhds_within), - have := ext_chart_at_target_mem_nhds_within' I x (hs hx'), - refine nhds_within_mono _ (inter_subset_right _ _) this } -end +cont_mdiff_on_iff_of_mem_maximal_atlas + (chart_mem_maximal_atlas I x) (chart_mem_maximal_atlas I' y) hs h2s /-- One can reformulate smoothness on a set as continuity on this set, and smoothness in any extended chart. -/ @@ -516,8 +505,8 @@ begin { simp only [w, hz] with mfld_simps }, { mfld_set_tac } }, { rintros ⟨hcont, hdiff⟩ x hx, - refine ((cont_diff_within_at_local_invariant_prop I I' n).lift_prop_within_at_iff $ - hcont x hx).mpr _, + refine (cont_diff_within_at_local_invariant_prop I I' n).lift_prop_within_at_iff.mpr _, + refine ⟨hcont x hx, _⟩, dsimp [cont_diff_within_at_prop], convert hdiff x (f x) (ext_chart_at I x x) (by simp only [hx] with mfld_simps) using 1, mfld_set_tac } @@ -673,10 +662,19 @@ end /-! ### Restriction to a smaller set -/ +lemma cont_mdiff_within_at.mono_of_mem (hf : cont_mdiff_within_at I I' n f s x) + (hts : s ∈ 𝓝[t] x) : cont_mdiff_within_at I I' n f t x := +structure_groupoid.local_invariant_prop.lift_prop_within_at_mono_of_mem + (cont_diff_within_at_prop_mono_of_mem I I' n) hf hts + lemma cont_mdiff_within_at.mono (hf : cont_mdiff_within_at I I' n f s x) (hts : t ⊆ s) : cont_mdiff_within_at I I' n f t x := -structure_groupoid.local_invariant_prop.lift_prop_within_at_mono - (cont_diff_within_at_prop_mono I I' n) hf hts +hf.mono_of_mem $ mem_of_superset self_mem_nhds_within hts + +lemma cont_mdiff_within_at_congr_nhds (hst : 𝓝[s] x = 𝓝[t] x) : + cont_mdiff_within_at I I' n f s x ↔ cont_mdiff_within_at I I' n f t x := +⟨λ h, h.mono_of_mem $ hst ▸ self_mem_nhds_within, + λ h, h.mono_of_mem $ hst.symm ▸ self_mem_nhds_within⟩ lemma cont_mdiff_at.cont_mdiff_within_at (hf : cont_mdiff_at I I' n f x) : cont_mdiff_within_at I I' n f s x := @@ -725,9 +723,17 @@ h.cont_mdiff_at hx include Is -lemma cont_mdiff_on_ext_chart_at : - cont_mdiff_on I 𝓘(𝕜, E) n (ext_chart_at I x) (chart_at H x).source := -λ x' hx', (cont_mdiff_at_ext_chart_at' hx').cont_mdiff_within_at +lemma cont_mdiff_on_iff_source_of_mem_maximal_atlas + (he : e ∈ maximal_atlas I M) (hs : s ⊆ e.source) : + cont_mdiff_on I I' n f s ↔ cont_mdiff_on 𝓘(𝕜, E) I' n (f ∘ (e.extend I).symm) (e.extend I '' s) := +begin + simp_rw [cont_mdiff_on, set.ball_image_iff], + refine forall₂_congr (λ x hx, _), + rw [cont_mdiff_within_at_iff_source_of_mem_maximal_atlas he (hs hx)], + apply cont_mdiff_within_at_congr_nhds, + simp_rw [nhds_within_eq_iff_eventually_eq, + e.extend_symm_preimage_inter_range_eventually_eq I hs (hs hx)] +end include I's @@ -1026,7 +1032,16 @@ end composition /-! ### Atlas members are smooth -/ section atlas -variables {e : local_homeomorph M H} +lemma cont_mdiff_model : cont_mdiff I 𝓘(𝕜, E) n I := +begin + intro x, + refine (cont_mdiff_at_iff _ _).mpr ⟨I.continuous_at, _⟩, + simp only with mfld_simps, + refine cont_diff_within_at_id.congr_of_eventually_eq _ _, + { exact eventually_eq_of_mem self_mem_nhds_within (λ x₂, I.right_inv) }, + simp_rw [function.comp_apply, I.left_inv, id_def] +end + include Is /-- An atlas member is `C^n` for any `n`. -/ @@ -1043,15 +1058,36 @@ cont_mdiff_on.of_le ((cont_diff_within_at_local_invariant_prop I I ∞).lift_prop_on_symm_of_mem_maximal_atlas (cont_diff_within_at_prop_id I) h) le_top +lemma cont_mdiff_at_of_mem_maximal_atlas (h : e ∈ maximal_atlas I M) (hx : x ∈ e.source) : + cont_mdiff_at I I n e x := +(cont_mdiff_on_of_mem_maximal_atlas h).cont_mdiff_at $ e.open_source.mem_nhds hx + +lemma cont_mdiff_at_symm_of_mem_maximal_atlas {x : H} + (h : e ∈ maximal_atlas I M) (hx : x ∈ e.target) : cont_mdiff_at I I n e.symm x := +(cont_mdiff_on_symm_of_mem_maximal_atlas h).cont_mdiff_at $ e.open_target.mem_nhds hx + lemma cont_mdiff_on_chart : cont_mdiff_on I I n (chart_at H x) (chart_at H x).source := -cont_mdiff_on_of_mem_maximal_atlas - ((cont_diff_groupoid ⊤ I).chart_mem_maximal_atlas x) +cont_mdiff_on_of_mem_maximal_atlas $ chart_mem_maximal_atlas I x lemma cont_mdiff_on_chart_symm : cont_mdiff_on I I n (chart_at H x).symm (chart_at H x).target := -cont_mdiff_on_symm_of_mem_maximal_atlas - ((cont_diff_groupoid ⊤ I).chart_mem_maximal_atlas x) +cont_mdiff_on_symm_of_mem_maximal_atlas $ chart_mem_maximal_atlas I x + +lemma cont_mdiff_at_extend {x : M} (he : e ∈ maximal_atlas I M) (hx : x ∈ e.source) : + cont_mdiff_at I 𝓘(𝕜, E) n (e.extend I) x := +(cont_mdiff_model _).comp x $ cont_mdiff_at_of_mem_maximal_atlas he hx + +lemma cont_mdiff_at_ext_chart_at' {x' : M} (h : x' ∈ (chart_at H x).source) : + cont_mdiff_at I 𝓘(𝕜, E) n (ext_chart_at I x) x' := +cont_mdiff_at_extend (chart_mem_maximal_atlas I x) h + +lemma cont_mdiff_at_ext_chart_at : cont_mdiff_at I 𝓘(𝕜, E) n (ext_chart_at I x) x := +cont_mdiff_at_ext_chart_at' $ mem_chart_source H x + +lemma cont_mdiff_on_ext_chart_at : + cont_mdiff_on I 𝓘(𝕜, E) n (ext_chart_at I x) (chart_at H x).source := +λ x' hx', (cont_mdiff_at_ext_chart_at' hx').cont_mdiff_within_at end atlas diff --git a/src/geometry/manifold/local_invariant_properties.lean b/src/geometry/manifold/local_invariant_properties.lean index f4ee2b421d5ad..e1e770aac570a 100644 --- a/src/geometry/manifold/local_invariant_properties.lean +++ b/src/geometry/manifold/local_invariant_properties.lean @@ -223,13 +223,13 @@ include hG /-- `lift_prop_within_at P f s x` is equivalent to a definition where we restrict the set we are considering to the domain of the charts at `x` and `f x`. -/ -lemma lift_prop_within_at_iff {f : M → M'} (hf : continuous_within_at f s x) : +lemma lift_prop_within_at_iff {f : M → M'} : lift_prop_within_at P f s x ↔ - P ((chart_at H' (f x)) ∘ f ∘ (chart_at H x).symm) + continuous_within_at f s x ∧ P ((chart_at H' (f x)) ∘ f ∘ (chart_at H x).symm) ((chart_at H x).target ∩ (chart_at H x).symm ⁻¹' (s ∩ f ⁻¹' (chart_at H' (f x)).source)) (chart_at H x x) := begin - rw [lift_prop_within_at, iff_true_intro hf, true_and, hG.congr_set], + refine and_congr_right (λ hf, hG.congr_set _), exact local_homeomorph.preimage_eventually_eq_target_inter_preimage_inter hf (mem_chart_source H x) (chart_source_mem_nhds H' (f x)) end @@ -443,15 +443,25 @@ lemma lift_prop_on_congr_iff (h₁ : ∀ y ∈ s, g' y = g y) : omit hG +lemma lift_prop_within_at_mono_of_mem + (mono_of_mem : ∀ ⦃s x t⦄ ⦃f : H → H'⦄, s ∈ 𝓝[t] x → P f s x → P f t x) + (h : lift_prop_within_at P g s x) (hst : s ∈ 𝓝[t] x) : + lift_prop_within_at P g t x := +begin + refine ⟨h.1.mono_of_mem hst, mono_of_mem _ h.2⟩, + simp_rw [← mem_map, (chart_at H x).symm.map_nhds_within_preimage_eq (mem_chart_target H x), + (chart_at H x).left_inv (mem_chart_source H x), hst] +end + lemma lift_prop_within_at_mono (mono : ∀ ⦃s x t⦄ ⦃f : H → H'⦄, t ⊆ s → P f s x → P f t x) - (h : lift_prop_within_at P g t x) (hst : s ⊆ t) : - lift_prop_within_at P g s x := + (h : lift_prop_within_at P g s x) (hts : t ⊆ s) : + lift_prop_within_at P g t x := begin - refine ⟨h.1.mono hst, _⟩, + refine ⟨h.1.mono hts, _⟩, apply mono (λ y hy, _) h.2, simp only with mfld_simps at hy, - simp only [hy, hst _] with mfld_simps, + simp only [hy, hts _] with mfld_simps, end lemma lift_prop_within_at_of_lift_prop_at diff --git a/src/geometry/manifold/smooth_manifold_with_corners.lean b/src/geometry/manifold/smooth_manifold_with_corners.lean index be02a60b98fc3..45cd85e25fc39 100644 --- a/src/geometry/manifold/smooth_manifold_with_corners.lean +++ b/src/geometry/manifold/smooth_manifold_with_corners.lean @@ -247,10 +247,17 @@ I.closed_embedding.closed_range lemma map_nhds_eq (x : H) : map I (𝓝 x) = 𝓝[range I] (I x) := I.closed_embedding.to_embedding.map_nhds_eq x +lemma map_nhds_within_eq (s : set H) (x : H) : map I (𝓝[s] x) = 𝓝[I '' s] (I x) := +I.closed_embedding.to_embedding.map_nhds_within_eq s x + lemma image_mem_nhds_within {x : H} {s : set H} (hs : s ∈ 𝓝 x) : I '' s ∈ 𝓝[range I] (I x) := I.map_nhds_eq x ▸ image_mem_map hs +lemma symm_map_nhds_within_image {x : H} {s : set H} : + map I.symm (𝓝[I '' s] (I x)) = 𝓝[s] x := +by rw [← I.map_nhds_within_eq, map_map, I.symm_comp_self, map_id] + lemma symm_map_nhds_within_range (x : H) : map I.symm (𝓝[range I] (I x)) = 𝓝 x := by rw [← I.map_nhds_eq, map_map, I.symm_comp_self, map_id] @@ -266,6 +273,18 @@ I.unique_diff_preimage e.open_source lemma unique_diff_at_image {x : H} : unique_diff_within_at 𝕜 (range I) (I x) := I.unique_diff _ (mem_range_self _) +lemma symm_continuous_within_at_comp_right_iff {X} [topological_space X] + {f : H → X} {s : set H} {x : H} : + continuous_within_at (f ∘ I.symm) (I.symm ⁻¹' s ∩ range I) (I x) ↔ continuous_within_at f s x := +begin + refine ⟨λ h, _, λ h, _⟩, + { have := h.comp I.continuous_within_at (maps_to_preimage _ _), + simp_rw [preimage_inter, preimage_preimage, I.left_inv, preimage_id', preimage_range, + inter_univ] at this, + rwa [function.comp.assoc, I.symm_comp_self] at this }, + { rw [← I.left_inv x] at h, exact h.comp I.continuous_within_at_symm (inter_subset_left _ _) } +end + protected lemma locally_compact [locally_compact_space E] (I : model_with_corners 𝕜 E H) : locally_compact_space H := begin @@ -779,6 +798,12 @@ lemma continuous_on_extend_symm : continuous_on (f.extend I).symm (f.extend I).target := λ y hy, (continuous_at_extend_symm' _ _ hy).continuous_within_at +lemma extend_symm_continuous_within_at_comp_right_iff {X} [topological_space X] {g : M → X} + {s : set M} {x : M} : + continuous_within_at (g ∘ (f.extend I).symm) ((f.extend I).symm ⁻¹' s ∩ range I) (f.extend I x) ↔ + continuous_within_at (g ∘ f.symm) (f.symm ⁻¹' s) (f x) := +by convert I.symm_continuous_within_at_comp_right_iff; refl + lemma is_open_extend_preimage' {s : set E} (hs : is_open s) : is_open ((f.extend I).source ∩ f.extend I ⁻¹' s) := (continuous_on_extend f I).preimage_open_of_open (is_open_extend_source _ _) hs From 67f362670ed961bcb80239dc40ca18bcd4289c77 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 6 Jan 2023 19:49:13 +0000 Subject: [PATCH 0163/1291] =?UTF-8?q?feat(data/set/function):=20`=E2=88=98?= =?UTF-8?q?`=20lemmas=20(#17819)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit A few lemmas about `inj_on`/`surj_on`/`bij_on` and `∘`. --- src/data/set/function.lean | 138 +++++++++++++++++++++++++++++++++++-- 1 file changed, 132 insertions(+), 6 deletions(-) diff --git a/src/data/set/function.lean b/src/data/set/function.lean index c23e7bd94023b..b51f2b6bb7f0e 100644 --- a/src/data/set/function.lean +++ b/src/data/set/function.lean @@ -95,7 +95,7 @@ lemma restrict_extend_range (f : α → β) (g : α → γ) (g' : β → γ) : by convert restrict_dite _ _ @[simp] lemma restrict_extend_compl_range (f : α → β) (g : α → γ) (g' : β → γ) : - (range f)ᶜ.restrict (extend f g g') = g' ∘ coe := + (range f)ᶜ.restrict (extend f g g') = g' ∘ coe := by convert restrict_dite_compl _ _ lemma range_extend_subset (f : α → β) (g : α → γ) (g' : β → γ) : @@ -134,7 +134,7 @@ by simp only [injective, subtype.ext_iff, coe_cod_restrict_apply] alias injective_cod_restrict ↔ _ _root_.function.injective.cod_restrict variables {s s₁ s₂ : set α} {t t₁ t₂ : set β} {p : set γ} {f f₁ f₂ f₃ : α → β} {g g₁ g₂ : β → γ} - {f' f₁' f₂' : β → α} {g' : γ → β} + {f' f₁' f₂' : β → α} {g' : γ → β} {a : α} {b : β} /-! ### Equality on a set -/ @@ -144,6 +144,7 @@ def eq_on (f₁ f₂ : α → β) (s : set α) : Prop := ∀ ⦃x⦄, x ∈ s → f₁ x = f₂ x @[simp] lemma eq_on_empty (f₁ f₂ : α → β) : eq_on f₁ f₂ ∅ := λ x, false.elim +@[simp] lemma eq_on_singleton : eq_on f₁ f₂ {a} ↔ f₁ a = f₂ a := by simp [set.eq_on] @[simp] lemma restrict_eq_restrict_iff : restrict s f₁ = restrict s f₂ ↔ eq_on f₁ f₂ s := restrict_eq_iff @@ -297,9 +298,8 @@ image_subset_iff.symm lemma maps_to.subset_preimage {f : α → β} {s : set α} {t : set β} (hf : maps_to f s t) : s ⊆ f ⁻¹' t := hf -@[simp] theorem maps_to_singleton {x : α} : maps_to f {x} t ↔ f x ∈ t := singleton_subset_iff - theorem maps_to_empty (f : α → β) (t : set β) : maps_to f ∅ t := empty_subset _ +@[simp] lemma maps_to_singleton : maps_to f {a} t ↔ f a ∈ t := singleton_subset_iff theorem maps_to.image_subset (h : maps_to f s t) : f '' s ⊆ t := maps_to'.1 h @@ -334,6 +334,13 @@ begin { simp [nat.iterate, ihn] } end +lemma maps_to_of_subsingleton' [subsingleton β] (f : α → β) (h : s.nonempty → t.nonempty) : + maps_to f s t := +λ a ha, subsingleton.mem_iff_nonempty.2 $ h ⟨a, ha⟩ + +lemma maps_to_of_subsingleton [subsingleton α] (f : α → α) (s : set α) : maps_to f s s := +maps_to_of_subsingleton' _ id + theorem maps_to.mono (hf : maps_to f s₁ t₁) (hs : s₂ ⊆ s₁) (ht : t₁ ⊆ t₂) : maps_to f s₂ t₂ := λ x hx, ht (hf $ hs hx) @@ -381,6 +388,12 @@ theorem maps_to_range (f : α → β) (s : set α) : maps_to f s (range f) := maps_to f (g '' s) t ↔ maps_to (f ∘ g) s t := ⟨λ h c hc, h ⟨c, hc, rfl⟩, λ h d ⟨c, hc⟩, hc.2 ▸ h hc.1⟩ +lemma maps_to.comp_left (g : β → γ) (hf : maps_to f s t) : maps_to (g ∘ f) s (g '' t) := +λ x hx, ⟨f x, hf hx, rfl⟩ + +lemma maps_to.comp_right {s : set β} {t : set γ} (hg : maps_to g s t) (f : α → β) : + maps_to (g ∘ f) (f ⁻¹' s) t := λ x hx, hg hx + @[simp] lemma maps_univ_to (f : α → β) (s : set β) : maps_to f univ s ↔ ∀ a, f a ∈ s := ⟨λ h a, h (mem_univ _), λ h x _, h x⟩ @@ -490,10 +503,20 @@ lemma inj_on_of_injective (h : injective f) (s : set α) : inj_on f s := alias inj_on_of_injective ← _root_.function.injective.inj_on +lemma inj_on_id (s : set α) : inj_on id s := injective_id.inj_on _ + theorem inj_on.comp (hg : inj_on g t) (hf: inj_on f s) (h : maps_to f s t) : inj_on (g ∘ f) s := λ x hx y hy heq, hf hx hy $ hg (h hx) (h hy) heq +lemma inj_on.iterate {f : α → α} {s : set α} (h : inj_on f s) (hf : maps_to f s s) : + ∀ n, inj_on (f^[n]) s +| 0 := inj_on_id _ +| (n + 1) := (inj_on.iterate n).comp h hf + +lemma inj_on_of_subsingleton [subsingleton α] (f : α → β) (s : set α) : inj_on f s := +(injective_of_subsingleton _).inj_on _ + lemma _root_.function.injective.inj_on_range (h : injective (g ∘ f)) : inj_on g (range f) := by { rintros _ ⟨x, rfl⟩ _ ⟨y, rfl⟩ H, exact congr_arg f (h H) } @@ -571,6 +594,8 @@ lemma surj_on_iff_exists_map_subtype : theorem surj_on_empty (f : α → β) (s : set α) : surj_on f s ∅ := empty_subset _ +@[simp] lemma surj_on_singleton : surj_on f s {b} ↔ b ∈ f '' s := singleton_subset_iff + theorem surj_on_image (f : α → β) (s : set α) : surj_on f s (f '' s) := subset.rfl theorem surj_on.comap_nonempty (h : surj_on f s t) (ht : t.nonempty) : s.nonempty := @@ -607,9 +632,29 @@ theorem surj_on.inter (h₁ : surj_on f s₁ t) (h₂ : surj_on f s₂ t) (h : i surj_on f (s₁ ∩ s₂) t := inter_self t ▸ h₁.inter_inter h₂ h +lemma surj_on_id (s : set α) : surj_on id s s := by simp [surj_on] + theorem surj_on.comp (hg : surj_on g t p) (hf : surj_on f s t) : surj_on (g ∘ f) s p := subset.trans hg $ subset.trans (image_subset g hf) $ (image_comp g f s) ▸ subset.refl _ +lemma surj_on.iterate {f : α → α} {s : set α} (h : surj_on f s s) : ∀ n, surj_on (f^[n]) s s +| 0 := surj_on_id _ +| (n + 1) := (surj_on.iterate n).comp h + +lemma surj_on.comp_left (hf : surj_on f s t) (g : β → γ) : surj_on (g ∘ f) s (g '' t) := +by { rw [surj_on, image_comp g f], exact image_subset _ hf } + +lemma surj_on.comp_right {s : set β} {t : set γ} (hf : surjective f) (hg : surj_on g s t) : + surj_on (g ∘ f) (f ⁻¹' s) t := +by rwa [surj_on, image_comp g f, image_preimage_eq _ hf] + +lemma surj_on_of_subsingleton' [subsingleton β] (f : α → β) (h : t.nonempty → s.nonempty) : + surj_on f s t := +λ a ha, subsingleton.mem_iff_nonempty.2 $ (h ⟨a, ha⟩).image _ + +lemma surj_on_of_subsingleton [subsingleton α] (f : α → α) (s : set α) : surj_on f s s := +surj_on_of_subsingleton' _ id + lemma surjective_iff_surj_on_univ : surjective f ↔ surj_on f univ univ := by simp [surjective, surj_on, subset_def] @@ -664,9 +709,11 @@ lemma bij_on.mk (h₁ : maps_to f s t) (h₂ : inj_on f s) (h₃ : surj_on f s t bij_on f s t := ⟨h₁, h₂, h₃⟩ -lemma bij_on_empty (f : α → β) : bij_on f ∅ ∅ := +@[simp] lemma bij_on_empty (f : α → β) : bij_on f ∅ ∅ := ⟨maps_to_empty f ∅, inj_on_empty f, surj_on_empty f ∅⟩ +@[simp] lemma bij_on_singleton : bij_on f {a} {b} ↔ f a = b := by simp [bij_on, eq_comm] + lemma bij_on.inter_maps_to (h₁ : bij_on f s₁ t₁) (h₂ : maps_to f s₂ t₂) (h₃ : s₁ ∩ f ⁻¹' t₂ ⊆ s₂) : bij_on f (s₁ ∩ s₂) (t₁ ∩ t₂) := ⟨h₁.maps_to.inter_inter h₂, h₁.inj_on.mono $ inter_subset_left _ _, @@ -703,10 +750,23 @@ lemma bij_on.image_eq (h : bij_on f s t) : f '' s = t := h.surj_on.image_eq_of_maps_to h.maps_to +lemma bij_on_id (s : set α) : bij_on id s s := ⟨s.maps_to_id, s.inj_on_id, s.surj_on_id⟩ + theorem bij_on.comp (hg : bij_on g t p) (hf : bij_on f s t) : bij_on (g ∘ f) s p := bij_on.mk (hg.maps_to.comp hf.maps_to) (hg.inj_on.comp hf.inj_on hf.maps_to) (hg.surj_on.comp hf.surj_on) +lemma bij_on.iterate {f : α → α} {s : set α} (h : bij_on f s s) : ∀ n, bij_on (f^[n]) s s +| 0 := s.bij_on_id +| (n + 1) := (bij_on.iterate n).comp h + +lemma bij_on_of_subsingleton' [subsingleton α] [subsingleton β] (f : α → β) + (h : s.nonempty ↔ t.nonempty) : bij_on f s t := +⟨maps_to_of_subsingleton' _ h.1, inj_on_of_subsingleton _ _, surj_on_of_subsingleton' _ h.2⟩ + +lemma bij_on_of_subsingleton [subsingleton α] (f : α → α) (s : set α) : bij_on f s s := +bij_on_of_subsingleton' _ iff.rfl + theorem bij_on.bijective (h : bij_on f s t) : bijective (h.maps_to.restrict f s t) := ⟨λ x y h', subtype.ext $ h.inj_on x.2 y.2 $ subtype.ext_iff.1 h', λ ⟨y, hy⟩, let ⟨x, hx, hxy⟩ := h.surj_on hy in ⟨⟨x, hx⟩, subtype.eq hxy⟩⟩ @@ -718,6 +778,8 @@ iff.intro (λ h, let ⟨map, inj, surj⟩ := h in ⟨iff.mpr injective_iff_inj_on_univ inj, iff.mpr surjective_iff_surj_on_univ surj⟩) +alias bijective_iff_bij_on_univ ↔ _root_.function.bijective.bij_on_univ _ + lemma bij_on.compl (hst : bij_on f s t) (hf : bijective f) : bij_on f sᶜ tᶜ := ⟨hst.surj_on.maps_to_compl hf.1, hf.1.inj_on _, hst.maps_to.surj_on_compl hf.2⟩ @@ -727,6 +789,9 @@ lemma bij_on.compl (hst : bij_on f s t) (hf : bijective f) : bij_on f sᶜ tᶜ def left_inv_on (f' : β → α) (f : α → β) (s : set α) : Prop := ∀ ⦃x⦄, x ∈ s → f' (f x) = x +@[simp] lemma left_inv_on_empty (f' : β → α) (f : α → β) : left_inv_on f' f ∅ := empty_subset _ +@[simp] lemma left_inv_on_singleton : left_inv_on f' f {a} ↔ f' (f a) = a := singleton_subset_iff + lemma left_inv_on.eq_on (h : left_inv_on f' f s) : eq_on (f' ∘ f) id s := h lemma left_inv_on.eq (h : left_inv_on f' f s) {x} (hx : x ∈ s) : f' (f x) = x := h hx @@ -752,6 +817,8 @@ theorem left_inv_on.surj_on (h : left_inv_on f' f s) (hf : maps_to f s t) : surj theorem left_inv_on.maps_to (h : left_inv_on f' f s) (hf : surj_on f s t) : maps_to f' t s := λ y hy, let ⟨x, hs, hx⟩ := hf hy in by rwa [← hx, h hs] +lemma left_inv_on_id (s : set α) : left_inv_on id id s := λ a _, rfl + theorem left_inv_on.comp (hf' : left_inv_on f' f s) (hg' : left_inv_on g' g t) (hf : maps_to f s t) : left_inv_on (f' ∘ g') (g ∘ f) s := @@ -793,6 +860,9 @@ theorem left_inv_on.image_image' (hf : left_inv_on f' f s) (hs : s₁ ⊆ s) : @[reducible] def right_inv_on (f' : β → α) (f : α → β) (t : set β) : Prop := left_inv_on f f' t +@[simp] lemma right_inv_on_empty (f' : β → α) (f : α → β) : right_inv_on f' f ∅ := empty_subset _ +@[simp] lemma right_inv_on_singleton : right_inv_on f' f {b} ↔ f (f' b) = b := singleton_subset_iff + lemma right_inv_on.eq_on (h : right_inv_on f' f t) : eq_on (f ∘ f') id t := h lemma right_inv_on.eq (h : right_inv_on f' f t) {y} (hy : y ∈ t) : f (f' y) = y := h hy @@ -815,6 +885,8 @@ hf.surj_on hf' theorem right_inv_on.maps_to (h : right_inv_on f' f t) (hf : surj_on f' t s) : maps_to f s t := h.maps_to hf +lemma right_inv_on_id (s : set α) : right_inv_on id id s := λ a _, rfl + theorem right_inv_on.comp (hf : right_inv_on f' f t) (hg : right_inv_on g' g p) (g'pt : maps_to g' p t) : right_inv_on (f' ∘ g') (g ∘ f) p := hg.comp hf g'pt @@ -843,8 +915,19 @@ theorem surj_on.left_inv_on_of_right_inv_on (hf : surj_on f s t) (hf' : right_in def inv_on (g : β → α) (f : α → β) (s : set α) (t : set β) : Prop := left_inv_on g f s ∧ right_inv_on g f t +@[simp] lemma inv_on_empty (f' : β → α) (f : α → β) : inv_on f' f ∅ ∅ := by simp [inv_on] +@[simp] lemma inv_on_singleton : inv_on f' f {a} {b} ↔ f' (f a) = a ∧ f (f' b) = b := +by simp [inv_on] + lemma inv_on.symm (h : inv_on f' f s t) : inv_on f f' t s := ⟨h.right, h.left⟩ +lemma inv_on_id (s : set α) : inv_on id id s s := ⟨s.left_inv_on_id, s.right_inv_on_id⟩ + +lemma inv_on.comp (hf : inv_on f' f s t) (hg : inv_on g' g t p) (fst : maps_to f s t) + (g'pt : maps_to g' p t) : + inv_on (f' ∘ g') (g ∘ f) s p := +⟨hf.1.comp hg.1 fst, hf.2.comp hg.2 g'pt⟩ + lemma inv_on.mono (h : inv_on f' f s t) (hs : s₁ ⊆ s) (ht : t₁ ⊆ t) : inv_on f' f s₁ t₁ := ⟨h.1.mono hs, h.2.mono ht⟩ @@ -855,6 +938,12 @@ theorem inv_on.bij_on (h : inv_on f' f s t) (hf : maps_to f s t) (hf' : maps_to bij_on f s t := ⟨hf, h.left.inj_on, h.right.surj_on hf'⟩ +lemma bij_on.symm {g : β → α} (h : inv_on f g t s) (hf : bij_on f s t) : bij_on g t s := +⟨h.2.maps_to hf.surj_on, h.1.inj_on, h.2.surj_on hf.maps_to⟩ + +lemma bij_on_comm {g : β → α} (h : inv_on f g t s) : bij_on f s t ↔ bij_on g t s := +⟨bij_on.symm h, bij_on.symm h.symm⟩ + end set /-! ### `inv_fun_on` is a left/right inverse -/ @@ -1144,6 +1233,8 @@ by simp end set +open set + lemma strict_mono_on.inj_on [linear_order α] [preorder β] {f : α → β} {s : set α} (H : strict_mono_on f s) : s.inj_on f := @@ -1311,7 +1402,7 @@ end function /-! ### Equivalences, permutations -/ namespace set -variables {p : β → Prop} [decidable_pred p] {f : α ≃ subtype p} {g : perm α} {s t : set α} +variables {p : β → Prop} [decidable_pred p] {f : α ≃ subtype p} {g g₁ g₂ : perm α} {s t : set α} protected lemma maps_to.extend_domain (h : maps_to g s t) : maps_to (g.extend_domain f) (coe ∘ f '' s) (coe ∘ f '' t) := @@ -1329,4 +1420,39 @@ protected lemma bij_on.extend_domain (h : set.bij_on g s t) : bij_on (g.extend_domain f) (coe ∘ f '' s) (coe ∘ f '' t) := ⟨h.maps_to.extend_domain, (g.extend_domain f).injective.inj_on _, h.surj_on.extend_domain⟩ +protected lemma left_inv_on.extend_domain (h : left_inv_on g₁ g₂ s) : + left_inv_on (g₁.extend_domain f) (g₂.extend_domain f) (coe ∘ f '' s) := +by { rintro _ ⟨a, ha, rfl⟩, simp_rw [extend_domain_apply_image, h ha] } + +protected lemma right_inv_on.extend_domain (h : right_inv_on g₁ g₂ t) : + right_inv_on (g₁.extend_domain f) (g₂.extend_domain f) (coe ∘ f '' t) := +by { rintro _ ⟨a, ha, rfl⟩, simp_rw [extend_domain_apply_image, h ha] } + +protected lemma inv_on.extend_domain (h : inv_on g₁ g₂ s t) : + inv_on (g₁.extend_domain f) (g₂.extend_domain f) (coe ∘ f '' s) (coe ∘ f '' t) := +⟨h.1.extend_domain, h.2.extend_domain⟩ + end set + +namespace equiv +variables (e : α ≃ β) {s : set α} {t : set β} + +lemma bij_on' (h₁ : maps_to e s t) (h₂ : maps_to e.symm t s) : bij_on e s t := +⟨h₁, e.injective.inj_on _, λ b hb, ⟨e.symm b, h₂ hb, apply_symm_apply _ _⟩⟩ + +protected lemma bij_on (h : ∀ a, e a ∈ t ↔ a ∈ s) : bij_on e s t := +e.bij_on' (λ a, (h _).2) $ λ b hb, (h _).1 $ by rwa apply_symm_apply + +lemma inv_on : inv_on e e.symm t s := +⟨e.right_inverse_symm.left_inv_on _, e.left_inverse_symm.left_inv_on _⟩ + +lemma bij_on_image : bij_on e s (e '' s) := (e.injective.inj_on _).bij_on_image +lemma bij_on_symm_image : bij_on e.symm (e '' s) s := e.bij_on_image.symm e.inv_on + +variables [decidable_eq α] {a b : α} + +lemma bij_on_swap (ha : a ∈ s) (hb : b ∈ s) : bij_on (swap a b) s s := +(swap a b).bij_on $ λ x, by obtain rfl | hxa := eq_or_ne x a; obtain rfl | hxb := eq_or_ne x b; + simp [*, swap_apply_of_ne_of_ne] + +end equiv From 117e93f82b5f959f8193857370109935291f0cc4 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 6 Jan 2023 22:53:08 +0000 Subject: [PATCH 0164/1291] feat(data/set/finite): Align `set.to_finset` and `set.finite.to_finset` (#17959) Match lemmas about `set.to_finset` and `set.finite.to_finset`. This mostly involves making sure we have all pairs of the form `set.to_finset_whatever`/`set.finite.to_finset_whatever`. Also add a few lemmas and tag existing ones with simp. --- src/algebra/big_operators/basic.lean | 2 +- src/analysis/analytic/inverse.lean | 8 +- src/analysis/locally_convex/weak_dual.lean | 2 +- src/combinatorics/simple_graph/basic.lean | 6 +- src/data/finset/basic.lean | 1 + src/data/finset/mul_antidiagonal.lean | 4 +- src/data/finset/pointwise.lean | 2 +- src/data/fintype/basic.lean | 85 ++++++++----- src/data/nat/nth.lean | 8 +- src/data/set/finite.lean | 119 +++++++++++++----- src/field_theory/polynomial_galois_group.lean | 7 +- src/probability/cond_count.lean | 3 +- src/topology/instances/ennreal.lean | 2 +- 13 files changed, 161 insertions(+), 88 deletions(-) diff --git a/src/algebra/big_operators/basic.lean b/src/algebra/big_operators/basic.lean index c6d418cdcec9b..07286aca97d6a 100644 --- a/src/algebra/big_operators/basic.lean +++ b/src/algebra/big_operators/basic.lean @@ -1679,7 +1679,7 @@ lemma disjoint_finset_sum_left {β : Type*} {i : finset β} {f : β → multiset multiset.disjoint (i.sum f) a ↔ ∀ b ∈ i, multiset.disjoint (f b) a := begin convert (@disjoint_sum_left _ a) (map f i.val), - simp [finset.mem_def, and.congr_left_iff, iff_self], + simp [and.congr_left_iff, iff_self], end lemma disjoint_finset_sum_right {β : Type*} {i : finset β} {f : β → multiset α} {a : multiset α} : diff --git a/src/analysis/analytic/inverse.lean b/src/analysis/analytic/inverse.lean index 91a9a365f5f3a..daf8eb3717456 100644 --- a/src/analysis/analytic/inverse.lean +++ b/src/analysis/analytic/inverse.lean @@ -126,7 +126,7 @@ begin ext k, simp [h] }, simp [formal_multilinear_series.comp, show n + 2 ≠ 1, by dec_trivial, A, finset.sum_union B, - apply_composition_ones, C, D], + apply_composition_ones, C, D, -set.to_finset_set_of], end /-! ### The right inverse of a formal multilinear series -/ @@ -196,7 +196,7 @@ begin = p 1 (λ (i : fin 1), q n v), { apply p.congr (composition.single_length hn) (λ j hj1 hj2, _), simp [apply_composition_single] }, - simp [formal_multilinear_series.comp, A, finset.sum_union B, C], + simp [formal_multilinear_series.comp, A, finset.sum_union B, C, -set.to_finset_set_of], end lemma comp_right_inv_aux2 @@ -229,7 +229,7 @@ begin continuous_linear_equiv.coe_apply, continuous_multilinear_curry_fin1_symm_apply] }, have N : 0 < n+2, by dec_trivial, simp [comp_right_inv_aux1 N, h, right_inv, lt_irrefl n, show n + 2 ≠ 1, by dec_trivial, - ← sub_eq_add_neg, sub_eq_zero, comp_right_inv_aux2], + ← sub_eq_add_neg, sub_eq_zero, comp_right_inv_aux2, -set.to_finset_set_of], end lemma right_inv_coeff (p : formal_multilinear_series 𝕜 E F) (i : E ≃L[𝕜] F) (n : ℕ) (hn : 2 ≤ n) : @@ -244,7 +244,7 @@ begin ext v, have N : 0 < n + 2, by dec_trivial, have : (p 1) (λ (i : fin 1), 0) = 0 := continuous_multilinear_map.map_zero _, - simp [comp_right_inv_aux1 N, lt_irrefl n, this, comp_right_inv_aux2] + simp [comp_right_inv_aux1 N, lt_irrefl n, this, comp_right_inv_aux2, -set.to_finset_set_of], end /-! ### Coincidence of the left and the right inverse -/ diff --git a/src/analysis/locally_convex/weak_dual.lean b/src/analysis/locally_convex/weak_dual.lean index df9abacbce362..39669e4384672 100644 --- a/src/analysis/locally_convex/weak_dual.lean +++ b/src/analysis/locally_convex/weak_dual.lean @@ -98,7 +98,7 @@ begin simp only [id.def], let U' := hU₁.to_finset, by_cases hU₃ : U.fst.nonempty, - { have hU₃' : U'.nonempty := hU₁.nonempty_to_finset.mpr hU₃, + { have hU₃' : U'.nonempty := hU₁.to_finset_nonempty.mpr hU₃, 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], diff --git a/src/combinatorics/simple_graph/basic.lean b/src/combinatorics/simple_graph/basic.lean index 593e33e9fb4b8..2434050c02421 100644 --- a/src/combinatorics/simple_graph/basic.lean +++ b/src/combinatorics/simple_graph/basic.lean @@ -960,7 +960,7 @@ lemma degree_compl [fintype (Gᶜ.neighbor_set v)] [fintype V] : begin classical, rw [← card_neighbor_set_union_compl_neighbor_set G v, set.to_finset_union], - simp [card_disjoint_union (set.to_finset_disjoint_iff.mpr (compl_neighbor_set_disjoint G v))], + simp [card_disjoint_union (set.disjoint_to_finset.mpr (compl_neighbor_set_disjoint G v))], end instance incidence_set_fintype [decidable_eq V] : fintype (G.incidence_set v) := @@ -1197,7 +1197,7 @@ begin { rw finset.insert_subset, split, { simpa, }, - { rw [neighbor_finset, set.to_finset_subset], + { rw [neighbor_finset, set.to_finset_subset_to_finset], exact G.common_neighbors_subset_neighbor_set_left _ _ } } end @@ -1207,7 +1207,7 @@ begin simp only [common_neighbors_top_eq, ← set.to_finset_card, set.to_finset_diff], rw finset.card_sdiff, { simp [finset.card_univ, h], }, - { simp only [set.to_finset_subset, set.subset_univ] }, + { simp only [set.to_finset_subset_to_finset, set.subset_univ] }, end end finite diff --git a/src/data/finset/basic.lean b/src/data/finset/basic.lean index 771c2eeb7e476..724ecea454c5e 100644 --- a/src/data/finset/basic.lean +++ b/src/data/finset/basic.lean @@ -158,6 +158,7 @@ instance has_decidable_eq [decidable_eq α] : decidable_eq (finset α) instance : has_mem α (finset α) := ⟨λ a s, a ∈ s.1⟩ theorem mem_def {a : α} {s : finset α} : a ∈ s ↔ a ∈ s.1 := iff.rfl +@[simp] lemma mem_val {a : α} {s : finset α} : a ∈ s.1 ↔ a ∈ s := iff.rfl @[simp] theorem mem_mk {a : α} {s nd} : a ∈ @finset.mk α s nd ↔ a ∈ s := iff.rfl diff --git a/src/data/finset/mul_antidiagonal.lean b/src/data/finset/mul_antidiagonal.lean index 154fbc72e8a24..701ae27716717 100644 --- a/src/data/finset/mul_antidiagonal.lean +++ b/src/data/finset/mul_antidiagonal.lean @@ -62,11 +62,11 @@ by simp [mul_antidiagonal, and_rotate] @[to_additive] lemma mul_antidiagonal_mono_left (h : u ⊆ s) : mul_antidiagonal hu ht a ⊆ mul_antidiagonal hs ht a := -set.finite.to_finset_subset.2 $ set.mul_antidiagonal_mono_left h +set.finite.to_finset_mono $ set.mul_antidiagonal_mono_left h @[to_additive] lemma mul_antidiagonal_mono_right (h : u ⊆ t) : mul_antidiagonal hs hu a ⊆ mul_antidiagonal hs ht a := -set.finite.to_finset_subset.2 $ set.mul_antidiagonal_mono_right h +set.finite.to_finset_mono $ set.mul_antidiagonal_mono_right h @[simp, to_additive] lemma swap_mem_mul_antidiagonal : x.swap ∈ finset.mul_antidiagonal hs ht a ↔ x ∈ finset.mul_antidiagonal ht hs a := diff --git a/src/data/finset/pointwise.lean b/src/data/finset/pointwise.lean index 2863687979276..b73271640aefe 100644 --- a/src/data/finset/pointwise.lean +++ b/src/data/finset/pointwise.lean @@ -1028,7 +1028,7 @@ t.zero_smul_subset.antisymm $ by simpa [mem_smul] using ht /-- A nonempty set is scaled by zero to the singleton set containing 0. -/ lemma zero_smul_finset {s : finset β} (h : s.nonempty) : (0 : α) • s = (0 : finset β) := -coe_injective $ by simpa using set.zero_smul_set h +coe_injective $ by simpa using @set.zero_smul_set α _ _ _ _ _ h lemma zero_smul_finset_subset (s : finset β) : (0 : α) • s ⊆ 0 := image_subset_iff.2 $ λ x _, mem_zero.2 $ zero_smul α x diff --git a/src/data/fintype/basic.lean b/src/data/fintype/basic.lean index 2d11dac574f2c..582ec043777c0 100644 --- a/src/data/fintype/basic.lean +++ b/src/data/fintype/basic.lean @@ -446,9 +446,6 @@ by cc @[simp] theorem mem_to_finset {s : set α} [fintype s] {a : α} : a ∈ s.to_finset ↔ a ∈ s := by simp [to_finset] -@[simp] theorem mem_to_finset_val {s : set α} [fintype s] {a : α} : a ∈ s.to_finset.1 ↔ a ∈ s := -mem_to_finset - /-- Many `fintype` instances for sets are defined using an extensionally equal `finset`. Rewriting `s.to_finset` with `set.to_finset_of_finset` replaces the term with such a `finset`. -/ theorem to_finset_of_finset {p : set α} (s : finset α) (H : ∀ x, x ∈ s ↔ x ∈ p) : @@ -472,51 +469,83 @@ by rw [←finset.coe_nonempty, coe_to_finset] s.to_finset = t.to_finset ↔ s = t := ⟨λ h, by rw [←s.coe_to_finset, h, t.coe_to_finset], λ h, by simp [h]; congr⟩ -@[simp, mono] lemma to_finset_subset [fintype s] [fintype t] : s.to_finset ⊆ t.to_finset ↔ s ⊆ t := +@[mono] +lemma to_finset_subset_to_finset [fintype s] [fintype t] : s.to_finset ⊆ t.to_finset ↔ s ⊆ t := by simp [finset.subset_iff, set.subset_def] -@[simp, mono] lemma to_finset_ssubset [fintype s] [fintype t] : s.to_finset ⊂ t.to_finset ↔ s ⊂ t := -by simp only [finset.ssubset_def, to_finset_subset, ssubset_def] +@[simp] lemma to_finset_ssubset [fintype s] {t : finset α} : s.to_finset ⊂ t ↔ s ⊂ t := +by rw [←finset.coe_ssubset, coe_to_finset] + +@[simp] lemma subset_to_finset {s : finset α} [fintype t] : s ⊆ t.to_finset ↔ ↑s ⊆ t := +by rw [←finset.coe_subset, coe_to_finset] + +@[simp] lemma ssubset_to_finset {s : finset α} [fintype t] : s ⊂ t.to_finset ↔ ↑s ⊂ t := +by rw [←finset.coe_ssubset, coe_to_finset] + +@[mono] +lemma to_finset_ssubset_to_finset [fintype s] [fintype t] : s.to_finset ⊂ t.to_finset ↔ s ⊂ t := +by simp only [finset.ssubset_def, to_finset_subset_to_finset, ssubset_def] + +@[simp] lemma to_finset_subset [fintype s] {t : finset α} : s.to_finset ⊆ t ↔ s ⊆ t := +by rw [←finset.coe_subset, coe_to_finset] -@[simp] theorem to_finset_disjoint_iff {s t : set α} [fintype s] [fintype t] : +alias to_finset_subset_to_finset ↔ _ to_finset_mono +alias to_finset_ssubset_to_finset ↔ _ to_finset_strict_mono + +@[simp] lemma disjoint_to_finset [fintype s] [fintype t] : disjoint s.to_finset t.to_finset ↔ disjoint s t := by simp only [←disjoint_coe, coe_to_finset] -@[simp] lemma to_finset_inter {α : Type*} [decidable_eq α] (s t : set α) [fintype (s ∩ t : set α)] - [fintype s] [fintype t] : (s ∩ t).to_finset = s.to_finset ∩ t.to_finset := -by { ext, simp } +section decidable_eq +variables [decidable_eq α] (s t) [fintype s] [fintype t] -@[simp] lemma to_finset_union {α : Type*} [decidable_eq α] (s t : set α) [fintype (s ∪ t : set α)] - [fintype s] [fintype t] : (s ∪ t).to_finset = s.to_finset ∪ t.to_finset := +@[simp] lemma to_finset_inter [fintype ↥(s ∩ t)] : (s ∩ t).to_finset = s.to_finset ∩ t.to_finset := by { ext, simp } -@[simp] lemma to_finset_diff {α : Type*} [decidable_eq α] (s t : set α) [fintype s] [fintype t] - [fintype (s \ t : set α)] : (s \ t).to_finset = s.to_finset \ t.to_finset := +@[simp] lemma to_finset_union [fintype ↥(s ∪ t)] : (s ∪ t).to_finset = s.to_finset ∪ t.to_finset := by { ext, simp } -lemma to_finset_ne_eq_erase {α : Type*} [decidable_eq α] [fintype α] (a : α) - [fintype {x : α | x ≠ a}] : {x : α | x ≠ a}.to_finset = finset.univ.erase a := +@[simp] lemma to_finset_diff [fintype ↥(s \ t)] : (s \ t).to_finset = s.to_finset \ t.to_finset := by { ext, simp } -@[simp] theorem to_finset_compl [decidable_eq α] [fintype α] (s : set α) [fintype s] [fintype ↥sᶜ] : - (sᶜ).to_finset = s.to_finsetᶜ := +@[simp] lemma to_finset_symm_diff [fintype ↥(s ∆ t)] : + (s ∆ t).to_finset = s.to_finset ∆ t.to_finset := +by { ext, simp [mem_symm_diff, finset.mem_symm_diff] } + +@[simp] lemma to_finset_compl [fintype α] [fintype ↥sᶜ] : sᶜ.to_finset = s.to_finsetᶜ := by { ext, simp } -@[simp] lemma to_finset_eq_univ [fintype α] {s : set α} [fintype s] : - s.to_finset = finset.univ ↔ s = set.univ := -by rw [← coe_inj, coe_to_finset, coe_univ] +end decidable_eq + +/- TODO The `↥` circumvents an elaboration bug. See comment on `set.to_finset_univ`. -/ +@[simp] lemma to_finset_empty [fintype ↥(∅ : set α)] : (∅ : set α).to_finset = ∅ := by { ext, simp } -/- TODO Without the coercion arrow (`↥`) there is an elaboration bug; +/- TODO Without the coercion arrow (`↥`) there is an elaboration bug in the following two; it essentially infers `fintype.{v} (set.univ.{u} : set α)` with `v` and `u` distinct. Reported in leanprover-community/lean#672 -/ -@[simp] lemma to_finset_univ [fintype ↥(set.univ : set α)] [fintype α] : +@[simp] lemma to_finset_univ [fintype α] [fintype ↥(set.univ : set α)] : (set.univ : set α).to_finset = finset.univ := -to_finset_eq_univ.2 rfl +by { ext, simp } + +@[simp] lemma to_finset_eq_empty [fintype s] : s.to_finset = ∅ ↔ s = ∅ := +by rw [←to_finset_empty, to_finset_inj] + +@[simp] lemma to_finset_eq_univ [fintype α] [fintype s] : s.to_finset = finset.univ ↔ s = univ := +by rw [← coe_inj, coe_to_finset, coe_univ] + +@[simp] lemma to_finset_set_of [fintype α] (p : α → Prop) [decidable_pred p] [fintype {x | p x}] : + {x | p x}.to_finset = finset.univ.filter p := +by { ext, simp } @[simp] lemma to_finset_ssubset_univ [fintype α] {s : set α} [fintype s] : s.to_finset ⊂ finset.univ ↔ s ⊂ univ := by rw [← coe_ssubset, coe_to_finset, coe_univ] +@[simp] +lemma to_finset_image [decidable_eq β] (f : α → β) (s : set α) [fintype s] [fintype (f '' s)] : + (f '' s).to_finset = s.to_finset.image f := +finset.coe_injective $ by simp + @[simp] lemma to_finset_range [decidable_eq α] [fintype β] (f : β → α) [fintype (set.range f)] : (set.range f).to_finset = finset.univ.image f := by { ext, simp } @@ -673,12 +702,6 @@ instance Prop.fintype : fintype Prop := instance subtype.fintype (p : α → Prop) [decidable_pred p] [fintype α] : fintype {x // p x} := fintype.subtype (univ.filter p) (by simp) -@[simp] lemma set.to_finset_eq_empty_iff {s : set α} [fintype s] : s.to_finset = ∅ ↔ s = ∅ := -by simp only [ext_iff, set.ext_iff, set.mem_to_finset, not_mem_empty, set.mem_empty_iff_false] - -@[simp] lemma set.to_finset_empty : (∅ : set α).to_finset = ∅ := -set.to_finset_eq_empty_iff.mpr rfl - /-- A set on a fintype, when coerced to a type, is a fintype. -/ def set_fintype [fintype α] (s : set α) [decidable_pred (∈ s)] : fintype s := subtype.fintype (λ x, x ∈ s) @@ -809,7 +832,7 @@ theorem quotient.fin_choice_eq {ι : Type*} [decidable_eq ι] [fintype ι] begin let q, swap, change quotient.lift_on q _ _ = _, have : q = ⟦λ i h, f i⟧, - { dsimp [q], + { dsimp only [q], exact quotient.induction_on (@finset.univ ι _).1 (λ l, quotient.fin_choice_aux_eq _ _) }, simp [this], exact setoid.refl _ diff --git a/src/data/nat/nth.lean b/src/data/nat/nth.lean index 524dc6eb1fec5..c1acd39cc0617 100644 --- a/src/data/nat/nth.lean +++ b/src/data/nat/nth.lean @@ -83,7 +83,7 @@ begin apply finset.card_erase_of_mem, rw [nth, set.finite.mem_to_finset], apply Inf_mem, - rwa [←hp''.nonempty_to_finset, ←finset.card_pos, hk], + rwa [←hp''.to_finset_nonempty, ←finset.card_pos, hk], end lemma nth_set_card {n : ℕ} (hp : (set_of p).finite) @@ -93,10 +93,10 @@ begin obtain hn | hn := le_or_lt n hp.to_finset.card, { exact nth_set_card_aux p hp _ hn }, rw nat.sub_eq_zero_of_le hn.le, - simp only [finset.card_eq_zero, set.finite_to_finset_eq_empty_iff, ←set.subset_empty_iff], + simp only [finset.card_eq_zero, set.finite.to_finset_eq_empty, ←set.subset_empty_iff], convert_to _ ⊆ {i : ℕ | p i ∧ ∀ (k : ℕ), k < hp.to_finset.card → nth p k < i}, { symmetry, - rw [←set.finite_to_finset_eq_empty_iff, ←finset.card_eq_zero, + rw [←set.finite.to_finset_eq_empty, ←finset.card_eq_zero, ←nat.sub_self hp.to_finset.card], { apply nth_set_card_aux p hp _ le_rfl }, { exact hp.subset (λ x hx, hx.1) } }, @@ -108,7 +108,7 @@ lemma nth_set_nonempty_of_lt_card {n : ℕ} (hp : (set_of p).finite) (hlt: n < h begin have hp': {i : ℕ | p i ∧ ∀ (k : ℕ), k < n → nth p k < i}.finite, { exact hp.subset (λ x hx, hx.1) }, - rw [←hp'.nonempty_to_finset, ←finset.card_pos, nth_set_card p hp], + rw [←hp'.to_finset_nonempty, ←finset.card_pos, nth_set_card p hp], exact nat.sub_pos_of_lt hlt, end diff --git a/src/data/set/finite.lean b/src/data/set/finite.lean index 641d6920551a0..0cd1144cbf2ce 100644 --- a/src/data/set/finite.lean +++ b/src/data/set/finite.lean @@ -88,11 +88,11 @@ This is the `fintype` projection for a `set.finite`. Note that because `finite` isn't a typeclass, this definition will not fire if it is made into an instance -/ -noncomputable def finite.fintype {s : set α} (h : s.finite) : fintype s := +protected noncomputable def finite.fintype {s : set α} (h : s.finite) : fintype s := h.nonempty_fintype.some /-- Using choice, get the `finset` that represents this `set.` -/ -noncomputable def finite.to_finset {s : set α} (h : s.finite) : finset α := +protected noncomputable def finite.to_finset {s : set α} (h : s.finite) : finset α := @set.to_finset _ _ h.fintype theorem finite.exists_finset {s : set α} (h : s.finite) : @@ -119,51 +119,102 @@ protected lemma finite_or_infinite (s : set α) : s.finite ∨ s.infinite := em /-! ### Basic properties of `set.finite.to_finset` -/ -section finite_to_finset -variables {s t : set α} +namespace finite +variables {s t : set α} {a : α} {hs : s.finite} {ht : t.finite} -@[simp] lemma finite.coe_to_finset {s : set α} (h : s.finite) : (h.to_finset : set α) = s := -@set.coe_to_finset _ s h.fintype - -@[simp] theorem finite.mem_to_finset {s : set α} (h : s.finite) {a : α} : a ∈ h.to_finset ↔ a ∈ s := +@[simp] protected lemma mem_to_finset (h : s.finite) : a ∈ h.to_finset ↔ a ∈ s := @mem_to_finset _ _ h.fintype _ -@[simp] theorem finite.nonempty_to_finset {s : set α} (h : s.finite) : - h.to_finset.nonempty ↔ s.nonempty := +@[simp] protected lemma coe_to_finset (h : s.finite) : (h.to_finset : set α) = s := +@coe_to_finset _ _ h.fintype + +@[simp] protected lemma to_finset_nonempty (h : s.finite) : h.to_finset.nonempty ↔ s.nonempty := by rw [← finset.coe_nonempty, finite.coe_to_finset] -@[simp] lemma finite.coe_sort_to_finset {s : set α} (h : s.finite) : - (h.to_finset : Type*) = s := +/-- Note that this is an equality of types not holding definitionally. Use wisely. -/ +lemma coe_sort_to_finset (h : s.finite) : ↥h.to_finset = ↥s := by rw [← finset.coe_sort_coe _, h.coe_to_finset] -@[simp] lemma finite_empty_to_finset (h : (∅ : set α).finite) : h.to_finset = ∅ := -by rw [← finset.coe_inj, h.coe_to_finset, finset.coe_empty] +@[simp] protected lemma to_finset_inj : hs.to_finset = ht.to_finset ↔ s = t := +@to_finset_inj _ _ _ hs.fintype ht.fintype -@[simp] lemma finite_univ_to_finset [fintype α] (h : (set.univ : set α).finite) : - h.to_finset = finset.univ := -finset.ext $ by simp +@[simp] lemma to_finset_subset {t : finset α} : hs.to_finset ⊆ t ↔ s ⊆ t := +by rw [←finset.coe_subset, finite.coe_to_finset] -@[simp] lemma finite.to_finset_inj {s t : set α} {hs : s.finite} {ht : t.finite} : - hs.to_finset = ht.to_finset ↔ s = t := -by simp only [←finset.coe_inj, finite.coe_to_finset] +@[simp] lemma to_finset_ssubset {t : finset α} : hs.to_finset ⊂ t ↔ s ⊂ t := +by rw [←finset.coe_ssubset, finite.coe_to_finset] -lemma subset_to_finset_iff {s : finset α} {t : set α} (ht : t.finite) : - s ⊆ ht.to_finset ↔ ↑s ⊆ t := -by rw [← finset.coe_subset, ht.coe_to_finset] +@[simp] lemma subset_to_finset {s : finset α} : s ⊆ ht.to_finset ↔ ↑s ⊆ t := +by rw [←finset.coe_subset, finite.coe_to_finset] -@[simp] lemma finite_to_finset_eq_empty_iff {s : set α} {h : s.finite} : - h.to_finset = ∅ ↔ s = ∅ := -by simp only [←finset.coe_inj, finite.coe_to_finset, finset.coe_empty] +@[simp] lemma ssubset_to_finset {s : finset α} : s ⊂ ht.to_finset ↔ ↑s ⊂ t := +by rw [←finset.coe_ssubset, finite.coe_to_finset] -@[simp, mono] lemma finite.to_finset_subset {hs : s.finite} {ht : t.finite} : - hs.to_finset ⊆ ht.to_finset ↔ s ⊆ t := +@[mono] protected lemma to_finset_subset_to_finset : hs.to_finset ⊆ ht.to_finset ↔ s ⊆ t := by simp only [← finset.coe_subset, finite.coe_to_finset] -@[simp, mono] lemma finite.to_finset_ssubset {hs : s.finite} {ht : t.finite} : - hs.to_finset ⊂ ht.to_finset ↔ s ⊂ t := +@[mono] protected lemma to_finset_ssubset_to_finset : hs.to_finset ⊂ ht.to_finset ↔ s ⊂ t := by simp only [← finset.coe_ssubset, finite.coe_to_finset] -end finite_to_finset +alias finite.to_finset_subset_to_finset ↔ _ to_finset_mono +alias finite.to_finset_ssubset_to_finset ↔ _ to_finset_strict_mono + +attribute [protected] to_finset_mono to_finset_strict_mono + +@[simp] protected lemma to_finset_set_of [fintype α] (p : α → Prop) [decidable_pred p] + (h : {x | p x}.finite) : + h.to_finset = finset.univ.filter p := +by { ext, simp } + +@[simp] lemma disjoint_to_finset {hs : s.finite} {ht : t.finite} : + disjoint hs.to_finset ht.to_finset ↔ disjoint s t := +@disjoint_to_finset _ _ _ hs.fintype ht.fintype + +protected lemma to_finset_inter [decidable_eq α] (hs : s.finite) (ht : t.finite) + (h : (s ∩ t).finite) : h.to_finset = hs.to_finset ∩ ht.to_finset := +by { ext, simp } + +protected lemma to_finset_union [decidable_eq α] (hs : s.finite) (ht : t.finite) + (h : (s ∪ t).finite) : h.to_finset = hs.to_finset ∪ ht.to_finset := +by { ext, simp } + +protected lemma to_finset_diff [decidable_eq α] (hs : s.finite) (ht : t.finite) + (h : (s \ t).finite) : h.to_finset = hs.to_finset \ ht.to_finset := +by { ext, simp } + +protected lemma to_finset_symm_diff [decidable_eq α] (hs : s.finite) (ht : t.finite) + (h : (s ∆ t).finite) : h.to_finset = hs.to_finset ∆ ht.to_finset := +by { ext, simp [mem_symm_diff, finset.mem_symm_diff] } + +protected lemma to_finset_compl [decidable_eq α] [fintype α] (hs : s.finite) (h : sᶜ.finite) : + h.to_finset = hs.to_finsetᶜ := +by { ext, simp } + +@[simp] protected lemma to_finset_empty (h : (∅ : set α).finite) : h.to_finset = ∅ := +by { ext, simp } + +-- Note: Not `simp` because `set.finite.to_finset_set_of` already proves it +@[simp] protected lemma to_finset_univ [fintype α] (h : (set.univ : set α).finite) : + h.to_finset = finset.univ := +by { ext, simp } + +@[simp] protected lemma to_finset_eq_empty {h : s.finite} : h.to_finset = ∅ ↔ s = ∅ := +@to_finset_eq_empty _ _ h.fintype + +@[simp] protected lemma to_finset_eq_univ [fintype α] {h : s.finite} : + h.to_finset = finset.univ ↔ s = univ := +@to_finset_eq_univ _ _ _ h.fintype + +protected lemma to_finset_image [decidable_eq β] (f : α → β) (hs : s.finite) (h : (f '' s).finite) : + h.to_finset = hs.to_finset.image f := +by { ext, simp } + +@[simp] protected lemma to_finset_range [decidable_eq α] [fintype β] (f : β → α) + (h : (range f).finite) : + h.to_finset = finset.univ.image f := +by { ext, simp } + +end finite /-! ### Fintype instances @@ -623,7 +674,7 @@ theorem exists_finite_iff_finset {p : set α → Prop} : /-- There are finitely many subsets of a given finite set -/ lemma finite.finite_subsets {α : Type u} {a : set α} (h : a.finite) : {b | b ⊆ a}.finite := ⟨fintype.of_finset ((finset.powerset h.to_finset).map finset.coe_emb.1) $ λ s, - by simpa [← @exists_finite_iff_finset α (λ t, t ⊆ a ∧ t = s), subset_to_finset_iff, + by simpa [← @exists_finite_iff_finset α (λ t, t ⊆ a ∧ t = s), finite.subset_to_finset, ← and.assoc] using h.subset⟩ /-- Finite product of finite sets is finite -/ @@ -865,8 +916,8 @@ lemma card_ne_eq [fintype α] (a : α) [fintype {x : α | x ≠ a}] : fintype.card {x : α | x ≠ a} = fintype.card α - 1 := begin haveI := classical.dec_eq α, - rw [←to_finset_card, to_finset_ne_eq_erase, finset.card_erase_of_mem (finset.mem_univ _), - finset.card_univ], + rw [←to_finset_card, to_finset_set_of, finset.filter_ne', + finset.card_erase_of_mem (finset.mem_univ _), finset.card_univ], end diff --git a/src/field_theory/polynomial_galois_group.lean b/src/field_theory/polynomial_galois_group.lean index 7d5cfea07ace0..c6806e3124acf 100644 --- a/src/field_theory/polynomial_galois_group.lean +++ b/src/field_theory/polynomial_galois_group.lean @@ -349,10 +349,9 @@ lemma card_complex_roots_eq_card_real_add_card_not_gal_inv (p : ℚ[X]) : (gal_action_hom p ℂ (restrict p ℂ (complex.conj_ae.restrict_scalars ℚ))).support.card := begin by_cases hp : p = 0, - { simp_rw [hp, root_set_zero, set.to_finset_eq_empty_iff.mpr rfl, finset.card_empty, zero_add], - refine eq.symm (le_zero_iff.mp ((finset.card_le_univ _).trans (le_of_eq _))), - simp_rw [hp, root_set_zero, fintype.card_eq_zero_iff], - apply_instance }, + { haveI : is_empty (p.root_set ℂ) := by { rw [hp, root_set_zero], apply_instance }, + simp_rw [(gal_action_hom p ℂ _).support.eq_empty_of_is_empty, hp, root_set_zero, + set.to_finset_empty, finset.card_empty] }, have inj : function.injective (is_scalar_tower.to_alg_hom ℚ ℝ ℂ) := (algebra_map ℝ ℂ).injective, rw [←finset.card_image_of_injective _ subtype.coe_injective, ←finset.card_image_of_injective _ inj], diff --git a/src/probability/cond_count.lean b/src/probability/cond_count.lean index 9787b2d47b859..09a8ab8c12782 100644 --- a/src/probability/cond_count.lean +++ b/src/probability/cond_count.lean @@ -127,8 +127,7 @@ begin suffices : s ∩ t = s, { exact this ▸ λ x hx, hx.2 }, rw ← @set.finite.to_finset_inj _ _ _ (hsf.inter_of_left _) hsf, - exact finset.eq_of_subset_of_card_le - (set.finite.to_finset_subset.2 (s.inter_subset_left t)) h.symm.le + exact finset.eq_of_subset_of_card_le (set.finite.to_finset_mono $ s.inter_subset_left t) h.ge, end lemma cond_count_eq_zero_iff (hs : s.finite) : diff --git a/src/topology/instances/ennreal.lean b/src/topology/instances/ennreal.lean index 3efe8bb66fd6a..ae2089995f426 100644 --- a/src/topology/instances/ennreal.lean +++ b/src/topology/instances/ennreal.lean @@ -1012,7 +1012,7 @@ begin have oops := (le_trans hi (le_tsum' (@ennreal.summable _ a) i)).trans tsum_le_c, rw h at oops, exact c_ne_top (le_antisymm le_top oops), }, - simp only [obs, finite_empty, finite_empty_to_finset, finset.card_empty, + simp only [obs, finite_empty, finite.to_finset_empty, finset.card_empty, algebra_map.coe_zero, zero_le', exists_true_left], }, have hf : {i : ι | ε ≤ a i}.finite, from ennreal.finite_const_le_of_tsum_ne_top From 3e00d81bdcbf77c8188bbd18f5524ddc3ed8cac6 Mon Sep 17 00:00:00 2001 From: Paul Lezeau <72892199+Paul-Lez@users.noreply.github.com> Date: Fri, 6 Jan 2023 22:53:09 +0000 Subject: [PATCH 0165/1291] chore({ring_theory, field_theory}/*): fix imports to make #18021 less messy (#18065) PR #18021 makes some changes to the theory of `minpoly`, which has the net effect that the structure of imports has to be changed for some files. This is a bit painful so I've decided to open a new PR instead of doing it in #18021. This PR also moves the following definitions from `ring_theory/adjoin_root.lean` to `field_theory/minpoly/gcd_monoid.lean`: - minpoly.to_adjoin.injective - minpoly.equiv_adjoin - algebra.adjoin.power_basis' - power_basis.of_gen_mem_adjoin' Co-authored-by: Paul Lezeau --- src/field_theory/is_alg_closed/basic.lean | 1 + src/field_theory/minpoly/basic.lean | 1 - src/field_theory/minpoly/field.lean | 2 +- src/field_theory/minpoly/gcd_monoid.lean | 59 +++++++++++++++++-- .../number_field/embeddings.lean | 3 +- src/ring_theory/adjoin_root.lean | 46 +-------------- src/ring_theory/is_adjoin_root.lean | 2 +- src/ring_theory/polynomial/selmer.lean | 1 + src/ring_theory/roots_of_unity.lean | 1 + 9 files changed, 63 insertions(+), 53 deletions(-) diff --git a/src/field_theory/is_alg_closed/basic.lean b/src/field_theory/is_alg_closed/basic.lean index 9783253d9ef7c..b23f594b1fee5 100644 --- a/src/field_theory/is_alg_closed/basic.lean +++ b/src/field_theory/is_alg_closed/basic.lean @@ -7,6 +7,7 @@ Authors: Kenny Lau import field_theory.perfect_closure import field_theory.separable import ring_theory.adjoin.field +import ring_theory.localization.integral /-! # Algebraically Closed Field diff --git a/src/field_theory/minpoly/basic.lean b/src/field_theory/minpoly/basic.lean index 529dc70c0c938..ec940ccbf05d3 100644 --- a/src/field_theory/minpoly/basic.lean +++ b/src/field_theory/minpoly/basic.lean @@ -5,7 +5,6 @@ Authors: Chris Hughes, Johan Commelin -/ import data.polynomial.field_division import ring_theory.integral_closure -import ring_theory.polynomial.gauss_lemma /-! # Minimal polynomials diff --git a/src/field_theory/minpoly/field.lean b/src/field_theory/minpoly/field.lean index 5db46495a2df4..f32ab9138dfcd 100644 --- a/src/field_theory/minpoly/field.lean +++ b/src/field_theory/minpoly/field.lean @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Riccardo Brasca, Johan Commelin -/ import data.polynomial.field_division -import ring_theory.integral_closure import field_theory.minpoly.basic +import ring_theory.algebraic /-! # Minimal polynomials on an algebra over a field diff --git a/src/field_theory/minpoly/gcd_monoid.lean b/src/field_theory/minpoly/gcd_monoid.lean index 5e56f553bc8a6..81719806b4065 100644 --- a/src/field_theory/minpoly/gcd_monoid.lean +++ b/src/field_theory/minpoly/gcd_monoid.lean @@ -4,9 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Riccardo Brasca -/ import data.polynomial.field_division -import ring_theory.integral_closure -import ring_theory.polynomial.gauss_lemma +import ring_theory.adjoin_root import field_theory.minpoly.field +import ring_theory.polynomial.gauss_lemma /-! # Minimal polynomials over a GCD monoid @@ -31,13 +31,13 @@ open polynomial set function minpoly namespace minpoly +variables {R S : Type*} [comm_ring R] [comm_ring S] [is_domain R] [is_domain S] [algebra R S] section gcd_domain -variables {R S : Type*} (K L : Type*) [comm_ring R] [is_domain R] [normalized_gcd_monoid R] - [field K] [comm_ring S] [is_domain S] [algebra R K] [is_fraction_ring R K] [algebra R S] [field L] - [algebra S L] [algebra K L] [algebra R L] [is_scalar_tower R K L] [is_scalar_tower R S L] - {s : S} (hs : is_integral R s) +variables (K L : Type*) [normalized_gcd_monoid R] [field K] [algebra R K] [is_fraction_ring R K] + [field L] [algebra S L] [algebra K L] [algebra R L] [is_scalar_tower R K L] + [is_scalar_tower R S L] {s : S} (hs : is_integral R s) include hs @@ -121,4 +121,51 @@ end end gcd_domain +section adjoin_root + +noncomputable theory + +open algebra polynomial adjoin_root + +variables {R} {x : S} [normalized_gcd_monoid R] [no_zero_smul_divisors R S] + +lemma to_adjoin.injective (hx : is_integral R x) : + function.injective (minpoly.to_adjoin R x) := +begin + refine (injective_iff_map_eq_zero _).2 (λ P₁ hP₁, _), + obtain ⟨P, hP⟩ := mk_surjective (minpoly.monic hx) P₁, + by_cases hPzero : P = 0, + { simpa [hPzero] using hP.symm }, + have hPcont : P.content ≠ 0 := λ h, hPzero (content_eq_zero_iff.1 h), + rw [← hP, minpoly.to_adjoin_apply', lift_hom_mk, ← subalgebra.coe_eq_zero, + aeval_subalgebra_coe, set_like.coe_mk, P.eq_C_content_mul_prim_part, aeval_mul, aeval_C] at hP₁, + replace hP₁ := eq_zero_of_ne_zero_of_mul_left_eq_zero + ((map_ne_zero_iff _ (no_zero_smul_divisors.algebra_map_injective R S)).2 hPcont) hP₁, + obtain ⟨Q, hQ⟩ := minpoly.gcd_domain_dvd hx P.is_primitive_prim_part.ne_zero hP₁, + rw [P.eq_C_content_mul_prim_part] at hP, + simpa [hQ] using hP.symm +end + +/-- The algebra isomorphism `adjoin_root (minpoly R x) ≃ₐ[R] adjoin R x` -/ +@[simps] def equiv_adjoin (hx : is_integral R x) : + adjoin_root (minpoly R x) ≃ₐ[R] adjoin R ({x} : set S) := +alg_equiv.of_bijective (minpoly.to_adjoin R x) + ⟨minpoly.to_adjoin.injective hx, minpoly.to_adjoin.surjective R x⟩ + +/-- The `power_basis` of `adjoin R {x}` given by `x`. See `algebra.adjoin.power_basis` for a version +over a field. -/ +@[simps] def _root_.algebra.adjoin.power_basis' (hx : _root_.is_integral R x) : + _root_.power_basis R (algebra.adjoin R ({x} : set S)) := +power_basis.map (adjoin_root.power_basis' (minpoly.monic hx)) (minpoly.equiv_adjoin hx) + +/-- The power basis given by `x` if `B.gen ∈ adjoin R {x}`. -/ +@[simps] noncomputable def _root_.power_basis.of_gen_mem_adjoin' (B : _root_.power_basis R S) + (hint : is_integral R x) (hx : B.gen ∈ adjoin R ({x} : set S)) : + _root_.power_basis R S := +(algebra.adjoin.power_basis' hint).map $ + (subalgebra.equiv_of_eq _ _ $ power_basis.adjoin_eq_top_of_gen_mem_adjoin hx).trans + subalgebra.top_equiv + +end adjoin_root + end minpoly diff --git a/src/number_theory/number_field/embeddings.lean b/src/number_theory/number_field/embeddings.lean index 23616031e974f..3e5e993c7678b 100644 --- a/src/number_theory/number_field/embeddings.lean +++ b/src/number_theory/number_field/embeddings.lean @@ -4,9 +4,10 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Alex J. Best, Xavier Roblot -/ -import number_theory.number_field.basic import analysis.complex.polynomial import data.complex.basic +import field_theory.minpoly.gcd_monoid +import number_theory.number_field.basic /-! # Embeddings of number fields diff --git a/src/ring_theory/adjoin_root.lean b/src/ring_theory/adjoin_root.lean index ae14a401494fd..08f9229b1d320 100644 --- a/src/ring_theory/adjoin_root.lean +++ b/src/ring_theory/adjoin_root.lean @@ -5,7 +5,7 @@ Authors: Mario Carneiro, Chris Hughes -/ import algebra.algebra.basic import data.polynomial.field_division -import field_theory.minpoly.gcd_monoid +import field_theory.minpoly.basic import ring_theory.adjoin.basic import ring_theory.finite_presentation import ring_theory.finite_type @@ -408,6 +408,8 @@ by rw [minpoly_power_basis_gen hf', hf.leading_coeff, inv_one, C.map_one, mul_on end power_basis +section equiv + section minpoly variables [comm_ring R] [comm_ring S] [algebra R S] (x : S) (R) @@ -415,7 +417,6 @@ variables [comm_ring R] [comm_ring S] [algebra R S] (x : S) (R) open algebra polynomial /-- The surjective algebra morphism `R[X]/(minpoly R x) → R[x]`. - If `R` is a GCD domain and `x` is integral, this is an isomorphism, see `adjoin_root.minpoly.equiv_adjoin`. -/ @[simps] def minpoly.to_adjoin : adjoin_root (minpoly R x) →ₐ[R] adjoin R ({x} : set S) := @@ -443,49 +444,8 @@ begin refine ⟨mk (minpoly R x) X, by simpa using h.symm⟩ end -variables {R} {x} [is_domain R] [normalized_gcd_monoid R] [is_domain S] [no_zero_smul_divisors R S] - -lemma minpoly.to_adjoin.injective (hx : is_integral R x) : - function.injective (minpoly.to_adjoin R x) := -begin - refine (injective_iff_map_eq_zero _).2 (λ P₁ hP₁, _), - obtain ⟨P, hP⟩ := mk_surjective (minpoly.monic hx) P₁, - by_cases hPzero : P = 0, - { simpa [hPzero] using hP.symm }, - have hPcont : P.content ≠ 0 := λ h, hPzero (content_eq_zero_iff.1 h), - rw [← hP, minpoly.to_adjoin_apply', lift_hom_mk, ← subalgebra.coe_eq_zero, - aeval_subalgebra_coe, set_like.coe_mk, P.eq_C_content_mul_prim_part, aeval_mul, aeval_C] at hP₁, - replace hP₁ := eq_zero_of_ne_zero_of_mul_left_eq_zero - ((map_ne_zero_iff _ (no_zero_smul_divisors.algebra_map_injective R S)).2 hPcont) hP₁, - obtain ⟨Q, hQ⟩ := minpoly.gcd_domain_dvd hx P.is_primitive_prim_part.ne_zero hP₁, - rw [P.eq_C_content_mul_prim_part] at hP, - simpa [hQ] using hP.symm -end - -/-- The algebra isomorphism `adjoin_root (minpoly R x) ≃ₐ[R] adjoin R x` -/ -@[simps] def minpoly.equiv_adjoin (hx : is_integral R x) : - adjoin_root (minpoly R x) ≃ₐ[R] adjoin R ({x} : set S) := -alg_equiv.of_bijective (minpoly.to_adjoin R x) - ⟨minpoly.to_adjoin.injective hx, minpoly.to_adjoin.surjective R x⟩ - -/-- The `power_basis` of `adjoin R {x}` given by `x`. See `algebra.adjoin.power_basis` for a version -over a field. -/ -@[simps] def _root_.algebra.adjoin.power_basis' (hx : _root_.is_integral R x) : - _root_.power_basis R (algebra.adjoin R ({x} : set S)) := -power_basis.map (adjoin_root.power_basis' (minpoly.monic hx)) (minpoly.equiv_adjoin hx) - -/-- The power basis given by `x` if `B.gen ∈ adjoin R {x}`. -/ -@[simps] noncomputable def _root_.power_basis.of_gen_mem_adjoin' (B : _root_.power_basis R S) - (hint : is_integral R x) (hx : B.gen ∈ adjoin R ({x} : set S)) : - _root_.power_basis R S := -(algebra.adjoin.power_basis' hint).map $ - (subalgebra.equiv_of_eq _ _ $ power_basis.adjoin_eq_top_of_gen_mem_adjoin hx).trans - subalgebra.top_equiv - end minpoly -section equiv - section is_domain variables [comm_ring R] [is_domain R] [comm_ring S] [is_domain S] [algebra R S] diff --git a/src/ring_theory/is_adjoin_root.lean b/src/ring_theory/is_adjoin_root.lean index 5535b5f32778a..a9d74b7173734 100644 --- a/src/ring_theory/is_adjoin_root.lean +++ b/src/ring_theory/is_adjoin_root.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Anne Baanen -/ import data.polynomial.algebra_map -import ring_theory.adjoin_root +import field_theory.minpoly.gcd_monoid import ring_theory.power_basis /-! diff --git a/src/ring_theory/polynomial/selmer.lean b/src/ring_theory/polynomial/selmer.lean index f8ba5e77ea676..4d449cb563af7 100644 --- a/src/ring_theory/polynomial/selmer.lean +++ b/src/ring_theory/polynomial/selmer.lean @@ -5,6 +5,7 @@ Authors: Thomas Browning -/ import data.polynomial.unit_trinomial +import ring_theory.polynomial.gauss_lemma import tactic.linear_combination /-! diff --git a/src/ring_theory/roots_of_unity.lean b/src/ring_theory/roots_of_unity.lean index 220abd647157c..d4ba7b344f82f 100644 --- a/src/ring_theory/roots_of_unity.lean +++ b/src/ring_theory/roots_of_unity.lean @@ -8,6 +8,7 @@ import algebra.char_p.two import algebra.ne_zero import data.polynomial.ring_division import field_theory.finite.basic +import field_theory.minpoly.gcd_monoid import field_theory.separable import group_theory.specific_groups.cyclic import number_theory.divisors From 14e84382905302e3091536c4dcabb5bb09d63d21 Mon Sep 17 00:00:00 2001 From: Eric Date: Sat, 7 Jan 2023 01:49:11 +0000 Subject: [PATCH 0166/1291] feat(analysis/convex): add lemma for composition of convex functions (#15111) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Add lemma for when the composition of two convex functions is convex. Closes #14902 Co-authored-by: Yaël Dillies --- src/analysis/convex/function.lean | 52 ++++++++++++++++++++++++++++--- 1 file changed, 48 insertions(+), 4 deletions(-) diff --git a/src/analysis/convex/function.lean b/src/analysis/convex/function.lean index 1a6e1c735104b..27795ab9b7fed 100644 --- a/src/analysis/convex/function.lean +++ b/src/analysis/convex/function.lean @@ -27,7 +27,7 @@ a convex set. open finset linear_map set open_locale big_operators classical convex pointwise -variables {𝕜 E F β ι : Type*} +variables {𝕜 E F α β ι : Type*} section ordered_semiring variables [ordered_semiring 𝕜] @@ -36,10 +36,10 @@ section add_comm_monoid variables [add_comm_monoid E] [add_comm_monoid F] section ordered_add_comm_monoid -variables [ordered_add_comm_monoid β] +variables [ordered_add_comm_monoid α] [ordered_add_comm_monoid β] section has_smul -variables (𝕜) [has_smul 𝕜 E] [has_smul 𝕜 β] (s : set E) (f : E → β) +variables (𝕜) [has_smul 𝕜 E] [has_smul 𝕜 α] [has_smul 𝕜 β] (s : set E) (f : E → β) {g : β → α} /-- Convexity of functions -/ def convex_on : Prop := @@ -101,8 +101,52 @@ lemma strict_concave_on.subset {t : set E} (hf : strict_concave_on 𝕜 t f) (hs strict_concave_on 𝕜 s f := ⟨hs, λ x hx y hy, hf.2 (hst hx) (hst hy)⟩ -end has_smul +lemma convex_on.comp (hg : convex_on 𝕜 (f '' s) g) (hf : convex_on 𝕜 s f) + (hg' : monotone_on g (f '' s)) : convex_on 𝕜 s (g ∘ f) := +⟨hf.1, λ x hx y hy a b ha hb hab, (hg' (mem_image_of_mem f $ hf.1 hx hy ha hb hab) + (hg.1 (mem_image_of_mem f hx) (mem_image_of_mem f hy) ha hb hab) $ hf.2 hx hy ha hb hab).trans $ + hg.2 (mem_image_of_mem f hx) (mem_image_of_mem f hy) ha hb hab⟩ + +lemma concave_on.comp (hg : concave_on 𝕜 (f '' s) g) (hf : concave_on 𝕜 s f) + (hg' : monotone_on g (f '' s)) : concave_on 𝕜 s (g ∘ f) := +⟨hf.1, λ x hx y hy a b ha hb hab, + (hg.2 (mem_image_of_mem f hx) (mem_image_of_mem f hy) ha hb hab).trans $ + hg' (hg.1 (mem_image_of_mem f hx) (mem_image_of_mem f hy) ha hb hab) + (mem_image_of_mem f $ hf.1 hx hy ha hb hab) $ hf.2 hx hy ha hb hab⟩ + +lemma convex_on.comp_concave_on (hg : convex_on 𝕜 (f '' s) g) (hf : concave_on 𝕜 s f) + (hg' : antitone_on g (f '' s)) : convex_on 𝕜 s (g ∘ f) := +hg.dual.comp hf hg' + +lemma concave_on.comp_convex_on (hg : concave_on 𝕜 (f '' s) g) (hf : convex_on 𝕜 s f) + (hg' : antitone_on g (f '' s)) : concave_on 𝕜 s (g ∘ f) := +hg.dual.comp hf hg' + +lemma strict_convex_on.comp (hg : strict_convex_on 𝕜 (f '' s) g) (hf : strict_convex_on 𝕜 s f) + (hg' : strict_mono_on g (f '' s)) (hf' : s.inj_on f) : strict_convex_on 𝕜 s (g ∘ f) := +⟨hf.1, λ x hx y hy hxy a b ha hb hab, (hg' (mem_image_of_mem f $ hf.1 hx hy ha.le hb.le hab) + (hg.1 (mem_image_of_mem f hx) (mem_image_of_mem f hy) ha.le hb.le hab) $ + hf.2 hx hy hxy ha hb hab).trans $ + hg.2 (mem_image_of_mem f hx) (mem_image_of_mem f hy) (mt (hf' hx hy) hxy) ha hb hab⟩ + +lemma strict_concave_on.comp (hg : strict_concave_on 𝕜 (f '' s) g) (hf : strict_concave_on 𝕜 s f) + (hg' : strict_mono_on g (f '' s)) (hf' : s.inj_on f) : strict_concave_on 𝕜 s (g ∘ f) := +⟨hf.1, λ x hx y hy hxy a b ha hb hab, + (hg.2 (mem_image_of_mem f hx) (mem_image_of_mem f hy) (mt (hf' hx hy) hxy) ha hb hab).trans $ + hg' (hg.1 (mem_image_of_mem f hx) (mem_image_of_mem f hy) ha.le hb.le hab) + (mem_image_of_mem f $ hf.1 hx hy ha.le hb.le hab) $ hf.2 hx hy hxy ha hb hab⟩ + +lemma strict_convex_on.comp_strict_concave_on (hg : strict_convex_on 𝕜 (f '' s) g) + (hf : strict_concave_on 𝕜 s f) (hg' : strict_anti_on g (f '' s)) (hf' : s.inj_on f) : + strict_convex_on 𝕜 s (g ∘ f) := +hg.dual.comp hf hg' hf' +lemma strict_concave_on.comp_strict_convex_on (hg : strict_concave_on 𝕜 (f '' s) g) + (hf : strict_convex_on 𝕜 s f) (hg' : strict_anti_on g (f '' s)) (hf' : s.inj_on f) : + strict_concave_on 𝕜 s (g ∘ f) := +hg.dual.comp hf hg' hf' + +end has_smul section distrib_mul_action variables [has_smul 𝕜 E] [distrib_mul_action 𝕜 β] {s : set E} {f g : E → β} From 134625f523e737f650a6ea7f0c82a6177e45e622 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sat, 7 Jan 2023 07:47:21 +0000 Subject: [PATCH 0167/1291] chore(algebra/continued_fractions): golf some proofs (#18082) Also use `option.bind` instead of `do` notation. --- .../computation/basic.lean | 4 +- .../computation/terminates_iff_rat.lean | 3 +- .../computation/translations.lean | 63 +++---------------- .../terminated_stable.lean | 32 +++------- 4 files changed, 20 insertions(+), 82 deletions(-) diff --git a/src/algebra/continued_fractions/computation/basic.lean b/src/algebra/continued_fractions/computation/basic.lean index a8a7cbf956561..89945ec471bc8 100644 --- a/src/algebra/continued_fractions/computation/basic.lean +++ b/src/algebra/continued_fractions/computation/basic.lean @@ -129,8 +129,8 @@ For example, let `(v : ℚ) := 3.4`. The process goes as follows: -/ protected def stream (v : K) : stream $ option (int_fract_pair K) | 0 := some (int_fract_pair.of v) -| (n + 1) := do ap_n ← stream n, - if ap_n.fr = 0 then none else int_fract_pair.of ap_n.fr⁻¹ +| (n + 1) := (stream n).bind $ λ ap_n, + if ap_n.fr = 0 then none else some (int_fract_pair.of ap_n.fr⁻¹) /-- Shows that `int_fract_pair.stream` has the sequence property, that is once we return `none` at diff --git a/src/algebra/continued_fractions/computation/terminates_iff_rat.lean b/src/algebra/continued_fractions/computation/terminates_iff_rat.lean index 662c06e92dac3..53e9232e5047a 100644 --- a/src/algebra/continued_fractions/computation/terminates_iff_rat.lean +++ b/src/algebra/continued_fractions/computation/terminates_iff_rat.lean @@ -186,8 +186,7 @@ begin rwa [stream_q_nth_eq] at IH, have : (fr : K)⁻¹ = ((fr⁻¹ : ℚ) : K), by norm_cast, have coe_of_fr := (coe_of_rat_eq this), - simp [int_fract_pair.stream, IH.symm, v_eq_q, stream_q_nth_eq, fr_ne_zero], - exact congr_arg some coe_of_fr } } } + simpa [int_fract_pair.stream, IH.symm, v_eq_q, stream_q_nth_eq, fr_ne_zero] } } } end lemma coe_stream_rat_eq : diff --git a/src/algebra/continued_fractions/computation/translations.lean b/src/algebra/continued_fractions/computation/translations.lean index 59c53e9b5afe3..a7eeeac76d51a 100644 --- a/src/algebra/continued_fractions/computation/translations.lean +++ b/src/algebra/continued_fractions/computation/translations.lean @@ -68,15 +68,8 @@ parts of a value in case of termination. lemma succ_nth_stream_eq_none_iff : int_fract_pair.stream v (n + 1) = none ↔ (int_fract_pair.stream v n = none ∨ ∃ ifp, int_fract_pair.stream v n = some ifp ∧ ifp.fr = 0) := begin - cases stream_nth_eq : (int_fract_pair.stream v n) with ifp, - case option.none : { simp [stream_nth_eq, int_fract_pair.stream] }, - case option.some : - { cases ifp with _ fr, - by_cases h : fr = 0, -- `finish [int_fract_pair.stream]` closes both goals - { simp [int_fract_pair.stream, h, stream_nth_eq] }, - { suffices : ¬ (int_fract_pair.of fr⁻¹: option $ int_fract_pair K) = none, - by simp [int_fract_pair.stream, h, stream_nth_eq, this], - exact λ h, option.no_confusion h } } + rw [int_fract_pair.stream], + cases int_fract_pair.stream v n; simp [imp_false] end /-- @@ -88,31 +81,7 @@ lemma succ_nth_stream_eq_some_iff {ifp_succ_n : int_fract_pair K} : ↔ ∃ (ifp_n : int_fract_pair K), int_fract_pair.stream v n = some ifp_n ∧ ifp_n.fr ≠ 0 ∧ int_fract_pair.of ifp_n.fr⁻¹ = ifp_succ_n := -begin - split, - { assume stream_succ_nth_eq, - have : int_fract_pair.stream v (n + 1) ≠ none, by simp [stream_succ_nth_eq], - have : ¬int_fract_pair.stream v n = none - ∧ ¬(∃ ifp, int_fract_pair.stream v n = some ifp ∧ ifp.fr = 0), by - { have not_none_not_fract_zero, - from (not_iff_not_of_iff succ_nth_stream_eq_none_iff).elim_left this, - exact (not_or_distrib.elim_left not_none_not_fract_zero) }, - cases this with stream_nth_ne_none nth_fr_ne_zero, - replace nth_fr_ne_zero : ∀ ifp, int_fract_pair.stream v n = some ifp → ifp.fr ≠ 0, by - simpa using nth_fr_ne_zero, - obtain ⟨ifp_n, stream_nth_eq⟩ : ∃ ifp_n, int_fract_pair.stream v n = some ifp_n, from - option.ne_none_iff_exists'.mp stream_nth_ne_none, - existsi ifp_n, - have ifp_n_fr_ne_zero : ifp_n.fr ≠ 0, from nth_fr_ne_zero ifp_n stream_nth_eq, - cases ifp_n with _ ifp_n_fr, - suffices : int_fract_pair.of ifp_n_fr⁻¹ = ifp_succ_n, - by simpa [stream_nth_eq, ifp_n_fr_ne_zero], - simp only [int_fract_pair.stream, stream_nth_eq, ifp_n_fr_ne_zero, option.some_bind, if_false] - at stream_succ_nth_eq, - injection stream_succ_nth_eq }, - { rintro ⟨⟨_⟩, ifp_n_props⟩, -- `finish [int_fract_pair.stream, ifp_n_props]` closes this goal - simpa only [int_fract_pair.stream, ifp_n_props, option.some_bind, if_false] } -end +by simp [int_fract_pair.stream, ite_eq_iff] lemma exists_succ_nth_stream_of_fr_zero {ifp_succ_n : int_fract_pair K} (stream_succ_nth_eq : int_fract_pair.stream v (n + 1) = some ifp_succ_n) @@ -121,19 +90,10 @@ lemma exists_succ_nth_stream_of_fr_zero {ifp_succ_n : int_fract_pair K} begin -- get the witness from `succ_nth_stream_eq_some_iff` and prove that it has the additional -- properties - rcases (succ_nth_stream_eq_some_iff.elim_left stream_succ_nth_eq) with - ⟨ifp_n, stream_nth_eq, nth_fr_ne_zero, _⟩, - existsi ifp_n, - cases ifp_n with _ ifp_n_fr, - suffices : ifp_n_fr⁻¹ = ⌊ifp_n_fr⁻¹⌋, by simpa [stream_nth_eq], - have : int_fract_pair.of ifp_n_fr⁻¹ = ifp_succ_n := h_right_right, - cases ifp_succ_n with _ ifp_succ_n_fr, - change ifp_succ_n_fr = 0 at succ_nth_fr_eq_zero, - have : int.fract ifp_n_fr⁻¹ = ifp_succ_n_fr, by injection this, - have : int.fract ifp_n_fr⁻¹ = 0, by rwa [succ_nth_fr_eq_zero] at this, - calc - ifp_n_fr⁻¹ = int.fract ifp_n_fr⁻¹ + ⌊ifp_n_fr⁻¹⌋ : by rw (int.fract_add_floor ifp_n_fr⁻¹) - ... = ⌊ifp_n_fr⁻¹⌋ : by simp [‹int.fract ifp_n_fr⁻¹ = 0›] + rcases (succ_nth_stream_eq_some_iff.mp stream_succ_nth_eq) with + ⟨ifp_n, seq_nth_eq, nth_fr_ne_zero, rfl⟩, + refine ⟨ifp_n, seq_nth_eq, _⟩, + simpa only [int_fract_pair.of, int.fract, sub_eq_zero] using succ_nth_fr_eq_zero end end int_fract_pair @@ -185,12 +145,7 @@ Let's first show how the termination of one sequence implies the termination of lemma of_terminated_at_iff_int_fract_pair_seq1_terminated_at : (of v).terminated_at n ↔ (int_fract_pair.seq1 v).snd.terminated_at n := -begin - rw [terminated_at_iff_s_none, of], - rcases (int_fract_pair.seq1 v) with ⟨head, ⟨st⟩⟩, - cases st_n_eq : st n; - simp [of, st_n_eq, seq.map, seq.nth, stream.map, seq.terminated_at, stream.nth] -end +option.map_eq_none lemma of_terminated_at_n_iff_succ_nth_int_fract_pair_stream_eq_none : (of v).terminated_at n ↔ int_fract_pair.stream v (n + 1) = none := @@ -241,7 +196,7 @@ lemma nth_of_eq_some_of_nth_int_fract_pair_stream_fr_ne_zero {ifp_n : int_fract_ (stream_nth_eq : int_fract_pair.stream v n = some ifp_n) (nth_fr_ne_zero : ifp_n.fr ≠ 0) : (of v).s.nth n = some ⟨1, (int_fract_pair.of ifp_n.fr⁻¹).b⟩ := have int_fract_pair.stream v (n + 1) = some (int_fract_pair.of ifp_n.fr⁻¹), by - { cases ifp_n, simp [int_fract_pair.stream, stream_nth_eq, nth_fr_ne_zero], refl }, + { cases ifp_n, simp [int_fract_pair.stream, stream_nth_eq, nth_fr_ne_zero] }, nth_of_eq_some_of_succ_nth_int_fract_pair_stream this end values diff --git a/src/algebra/continued_fractions/terminated_stable.lean b/src/algebra/continued_fractions/terminated_stable.lean index 2e593c632282f..0bdcae77dd778 100644 --- a/src/algebra/continued_fractions/terminated_stable.lean +++ b/src/algebra/continued_fractions/terminated_stable.lean @@ -28,22 +28,14 @@ lemma continuants_aux_stable_step_of_terminated (terminated_at_n : g.terminated_ by { rw [terminated_at_iff_s_none] at terminated_at_n, simp only [terminated_at_n, continuants_aux] } -lemma continuants_aux_stable_of_terminated (succ_n_le_m : (n + 1) ≤ m) +lemma continuants_aux_stable_of_terminated (n_lt_m : n < m) (terminated_at_n : g.terminated_at n) : g.continuants_aux m = g.continuants_aux (n + 1) := begin - induction succ_n_le_m with m succ_n_le_m IH, - { refl }, - { have : g.continuants_aux (m + 1) = g.continuants_aux m, by - { have : n ≤ m - 1, from nat.le_pred_of_lt succ_n_le_m, - have : g.terminated_at (m - 1), from terminated_stable this terminated_at_n, - have stable_step : g.continuants_aux (m - 1 + 2) = g.continuants_aux (m - 1 + 1), from - continuants_aux_stable_step_of_terminated this, - have one_le_m : 1 ≤ m, from nat.one_le_of_lt succ_n_le_m, - have : m - 1 + 2 = m + 2 - 1, from tsub_add_eq_add_tsub one_le_m, - have : m - 1 + 1 = m + 1 - 1, from tsub_add_eq_add_tsub one_le_m, - simpa [*] using stable_step }, - exact (eq.trans this IH) } + refine nat.le_induction rfl (λ k hnk hk, _) _ n_lt_m, + rcases nat.exists_eq_add_of_lt hnk with ⟨k, rfl⟩, + refine (continuants_aux_stable_step_of_terminated _).trans hk, + exact terminated_stable (nat.le_add_right _ _) terminated_at_n end lemma convergents'_aux_stable_step_of_terminated {s : seq $ pair K} @@ -67,18 +59,10 @@ lemma convergents'_aux_stable_of_terminated (terminated_at_n : s.terminated_at n) : convergents'_aux s m = convergents'_aux s n := begin - induction n_le_m with m n_le_m IH generalizing s, + induction n_le_m with m n_le_m IH, { refl }, - { cases s_head_eq : s.head with gp_head, - case option.none { cases n; simp only [convergents'_aux, s_head_eq] }, - case option.some - { have : convergents'_aux s (n + 1) = convergents'_aux s n, from - convergents'_aux_stable_step_of_terminated terminated_at_n, - rw [←this], - have : s.tail.terminated_at n, by - simpa only [seq.terminated_at, seq.nth_tail] using (s.le_stable n.le_succ terminated_at_n), - have : convergents'_aux s.tail m = convergents'_aux s.tail n, from IH this, - simp only [convergents'_aux, s_head_eq, this] } } + { refine (convergents'_aux_stable_step_of_terminated _).trans IH, + exact s.terminated_stable n_le_m terminated_at_n } end lemma continuants_stable_of_terminated (n_le_m : n ≤ m) (terminated_at_n : g.terminated_at n) : From 10708587e81b68c763fcdb7505f279d52e569768 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Sat, 7 Jan 2023 08:55:23 +0000 Subject: [PATCH 0168/1291] chore(*): add mathlib4 synchronization comments (#18078) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.order.ring.with_top` * `data.bool.all_any` * `data.list.forall2` * `data.list.sections` * `deprecated.group` * `deprecated.ring` --- src/algebra/order/ring/with_top.lean | 3 +++ src/data/bool/all_any.lean | 3 +++ src/data/list/forall2.lean | 3 +++ src/data/list/sections.lean | 3 +++ src/deprecated/group.lean | 3 +++ src/deprecated/ring.lean | 3 +++ 6 files changed, 18 insertions(+) diff --git a/src/algebra/order/ring/with_top.lean b/src/algebra/order/ring/with_top.lean index 5483f53ec6f3f..dc036fafb540e 100644 --- a/src/algebra/order/ring/with_top.lean +++ b/src/algebra/order/ring/with_top.lean @@ -9,6 +9,9 @@ import algebra.order.ring.canonical /-! # Structures involving `*` and `0` on `with_top` and `with_bot` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The main results of this section are `with_top.canonically_ordered_comm_semiring` and `with_bot.comm_monoid_with_zero`. -/ diff --git a/src/data/bool/all_any.lean b/src/data/bool/all_any.lean index 2b7b3bbe292a7..150e38cf3ab26 100644 --- a/src/data/bool/all_any.lean +++ b/src/data/bool/all_any.lean @@ -8,6 +8,9 @@ import data.list.basic /-! # Boolean quantifiers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This proves a few properties about `list.all` and `list.any`, which are the `bool` universal and existential quantifiers. Their definitions are in core Lean. -/ diff --git a/src/data/list/forall2.lean b/src/data/list/forall2.lean index 4eea34cfad2d6..ad1040cbe08a9 100644 --- a/src/data/list/forall2.lean +++ b/src/data/list/forall2.lean @@ -8,6 +8,9 @@ import data.list.infix /-! # Double universal quantification on a list +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file provides an API for `list.forall₂` (definition in `data.list.defs`). `forall₂ R l₁ l₂` means that `l₁` and `l₂` have the same length, and whenever `a` is the nth element of `l₁`, and `b` is the nth element of `l₂`, then `R a b` is satisfied. diff --git a/src/data/list/sections.lean b/src/data/list/sections.lean index a389b765a290c..6ee2bb5893c57 100644 --- a/src/data/list/sections.lean +++ b/src/data/list/sections.lean @@ -8,6 +8,9 @@ import data.list.forall2 /-! # List sections +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves some stuff about `list.sections` (definition in `data.list.defs`). A section of a list of lists `[l₁, ..., lₙ]` is a list whose `i`-th element comes from the `i`-th list. -/ diff --git a/src/deprecated/group.lean b/src/deprecated/group.lean index c85156a5faca6..d6cdaa34ce16f 100644 --- a/src/deprecated/group.lean +++ b/src/deprecated/group.lean @@ -11,6 +11,9 @@ import algebra.hom.units /-! # Unbundled monoid and group homomorphisms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file is deprecated, and is no longer imported by anything in mathlib other than other deprecated files, and test files. You should not need to import it. diff --git a/src/deprecated/ring.lean b/src/deprecated/ring.lean index 9c9cbdda46206..315e6ff949eaf 100644 --- a/src/deprecated/ring.lean +++ b/src/deprecated/ring.lean @@ -8,6 +8,9 @@ import deprecated.group /-! # Unbundled semiring and ring homomorphisms (deprecated) +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file is deprecated, and is no longer imported by anything in mathlib other than other deprecated files, and test files. You should not need to import it. From d6e84a0d3db8910c99b3aa0c56be88fa8bab6f80 Mon Sep 17 00:00:00 2001 From: Kyle Miller Date: Sat, 7 Jan 2023 11:52:43 +0000 Subject: [PATCH 0169/1291] feat(combinatorics/simple_graph/connectivity): `walk.to_subgraph` (#17325) A construction for the subgraph consisting of the vertices and edges of a given walk. --- .../simple_graph/connectivity.lean | 74 +++++++++++++++++++ src/combinatorics/simple_graph/subgraph.lean | 29 ++++++++ 2 files changed, 103 insertions(+) diff --git a/src/combinatorics/simple_graph/connectivity.lean b/src/combinatorics/simple_graph/connectivity.lean index 84f21619b2cdb..4a445bd6f9bdd 100644 --- a/src/combinatorics/simple_graph/connectivity.lean +++ b/src/combinatorics/simple_graph/connectivity.lean @@ -337,6 +337,9 @@ by cases p; simp @[simp] lemma end_mem_support {u v : V} (p : G.walk u v) : v ∈ p.support := by induction p; simp [*] +@[simp] lemma support_nonempty {u v : V} (p : G.walk u v) : {w | w ∈ p.support}.nonempty := +⟨u, by simp⟩ + lemma mem_support_iff {u v w : V} (p : G.walk u v) : w ∈ p.support ↔ w = u ∨ w ∈ p.support.tail := by cases p; simp @@ -1513,6 +1516,77 @@ by { rw ← set.nonempty_iff_univ_nonempty, exact hconn u v } lemma connected.set_univ_walk_nonempty (hconn : G.connected) (u v : V) : (set.univ : set (G.walk u v)).nonempty := hconn.preconnected.set_univ_walk_nonempty u v +/-! ### Walks as subgraphs -/ + +namespace walk +variables {G G'} {u v w : V} + +/-- The subgraph consisting of the vertices and edges of the walk. -/ +@[simp] protected def to_subgraph : Π {u v : V}, G.walk u v → G.subgraph +| u _ nil := G.singleton_subgraph u +| _ _ (cons h p) := G.subgraph_of_adj h ⊔ p.to_subgraph + +lemma to_subgraph_cons_nil_eq_subgraph_of_adj (h : G.adj u v) : + (cons h nil).to_subgraph = G.subgraph_of_adj h := +by simp + +lemma mem_verts_to_subgraph (p : G.walk u v) : + w ∈ p.to_subgraph.verts ↔ w ∈ p.support := +begin + induction p with _ x y z h p' ih, + { simp }, + { have : w = y ∨ w ∈ p'.support ↔ w ∈ p'.support := + ⟨by rintro (rfl | h); simp [*], by simp { contextual := tt}⟩, + simp [ih, or_assoc, this] } +end + +@[simp] lemma verts_to_subgraph (p : G.walk u v) : p.to_subgraph.verts = {w | w ∈ p.support} := +set.ext (λ _, p.mem_verts_to_subgraph) + +lemma mem_edges_to_subgraph (p : G.walk u v) {e : sym2 V} : + e ∈ p.to_subgraph.edge_set ↔ e ∈ p.edges := +by induction p; simp [*] + +@[simp] lemma edge_set_to_subgraph (p : G.walk u v) : p.to_subgraph.edge_set = {e | e ∈ p.edges} := +set.ext (λ _, p.mem_edges_to_subgraph) + +@[simp] lemma to_subgraph_append (p : G.walk u v) (q : G.walk v w) : + (p.append q).to_subgraph = p.to_subgraph ⊔ q.to_subgraph := +by induction p; simp [*, sup_assoc] + +@[simp] lemma to_subgraph_reverse (p : G.walk u v) : + p.reverse.to_subgraph = p.to_subgraph := +begin + induction p, + { simp }, + { simp only [*, walk.to_subgraph, reverse_cons, to_subgraph_append, subgraph_of_adj_symm], + rw [sup_comm], + congr, + ext; simp [-set.bot_eq_empty], } +end + +@[simp] lemma to_subgraph_rotate [decidable_eq V] (c : G.walk v v) (h : u ∈ c.support) : + (c.rotate h).to_subgraph = c.to_subgraph := +by rw [rotate, to_subgraph_append, sup_comm, ← to_subgraph_append, take_spec] + +@[simp] lemma to_subgraph_map (f : G →g G') (p : G.walk u v) : + (p.map f).to_subgraph = p.to_subgraph.map f := +by induction p; simp [*, subgraph.map_sup] + +@[simp] lemma finite_neighbor_set_to_subgraph (p : G.walk u v) : + (p.to_subgraph.neighbor_set w).finite := +begin + induction p, + { rw [walk.to_subgraph, neighbor_set_singleton_subgraph], + apply set.to_finite, }, + { rw [walk.to_subgraph, subgraph.neighbor_set_sup], + refine set.finite.union _ p_ih, + refine set.finite.subset _ (neighbor_set_subgraph_of_adj_subset p_h), + apply set.to_finite, }, +end + +end walk + /-! ### Walks of a given length -/ section walk_counting diff --git a/src/combinatorics/simple_graph/subgraph.lean b/src/combinatorics/simple_graph/subgraph.lean index 4895851c4e9fd..e1a67790ddd9b 100644 --- a/src/combinatorics/simple_graph/subgraph.lean +++ b/src/combinatorics/simple_graph/subgraph.lean @@ -308,6 +308,18 @@ instance : bounded_order (subgraph G) := @[simp] lemma sup_adj {H₁ H₂ : subgraph G} {v w : V} : (H₁ ⊔ H₂).adj v w ↔ H₁.adj v w ∨ H₂.adj v w := iff.rfl +@[simp] lemma verts_sup {H H' : G.subgraph} : (H ⊔ H').verts = H.verts ∪ H'.verts := rfl + +@[simp] lemma verts_inf {H H' : G.subgraph} : (H ⊓ H').verts = H.verts ∩ H'.verts := rfl + +lemma neighbor_set_sup {H H' : G.subgraph} (v : V) : + (H ⊔ H').neighbor_set v = H.neighbor_set v ∪ H'.neighbor_set v := +by { ext w, simp } + +lemma neighbor_set_inf {H H' : G.subgraph} (v : V) : + (H ⊓ H').neighbor_set v = H.neighbor_set v ∩ H'.neighbor_set v := +by { ext w, simp } + @[simp] lemma edge_set_top : (⊤ : subgraph G).edge_set = G.edge_set := rfl @[simp] lemma edge_set_bot : (⊥ : subgraph G).edge_set = ∅ := @@ -386,6 +398,23 @@ begin exact ⟨_, _, h.2 ha, rfl, rfl⟩ } end +lemma map_sup {G : simple_graph V} {G' : simple_graph W} (f : G →g G') + {H H' : G.subgraph} : + (H ⊔ H').map f = H.map f ⊔ H'.map f := +begin + ext1, + { simp only [set.image_union, map_verts, verts_sup]}, + { ext, + simp only [relation.map, map_adj, sup_adj], + split, + { rintro ⟨a, b, h|h, rfl, rfl⟩, + { exact or.inl ⟨_, _, h, rfl, rfl⟩ }, + { exact or.inr ⟨_, _, h, rfl, rfl⟩ } }, + { rintro (⟨a, b, h, rfl, rfl⟩|⟨a, b, h, rfl, rfl⟩), + { exact ⟨_, _, or.inl h, rfl, rfl⟩ }, + { exact ⟨_, _, or.inr h, rfl, rfl⟩ } } }, +end + /-- Graph homomorphisms induce a contravariant function on subgraphs. -/ @[simps] protected def comap {G' : simple_graph W} (f : G →g G') (H : G'.subgraph) : G.subgraph := From ad1c83e8f956cfa2c7e0b805f468eb39420ac572 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Martin=20Dvo=C5=99=C3=A1k?= Date: Sat, 7 Jan 2023 11:52:45 +0000 Subject: [PATCH 0170/1291] feat(data/real/ennreal): `mul_div_mul_left` (#18051) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Rémi Bottinelli Co-authored-by: Floris van Doorn Co-authored-by: Yaël Dillies Co-authored-by: madvorak Co-authored-by: Yaël Dillies --- src/data/real/ennreal.lean | 10 ++++++++++ 1 file changed, 10 insertions(+) diff --git a/src/data/real/ennreal.lean b/src/data/real/ennreal.lean index 1f1eacdba696b..a46143ac853e7 100644 --- a/src/data/real/ennreal.lean +++ b/src/data/real/ennreal.lean @@ -1090,6 +1090,16 @@ begin simp [h'a, h'b], end +protected lemma mul_div_mul_left (a b : ℝ≥0∞) (hc : c ≠ 0) (hc' : c ≠ ⊤) : + c * a / (c * b) = a / b := +by rw [div_eq_mul_inv, div_eq_mul_inv, mul_inv (or.inl hc) (or.inl hc'), mul_mul_mul_comm, + ennreal.mul_inv_cancel hc hc', one_mul] + +protected lemma mul_div_mul_right (a b : ℝ≥0∞) (hc : c ≠ 0) (hc' : c ≠ ⊤) : + a * c / (b * c) = a / b := +by rw [div_eq_mul_inv, div_eq_mul_inv, mul_inv (or.inr hc') (or.inr hc), mul_mul_mul_comm, + ennreal.mul_inv_cancel hc hc', mul_one] + protected lemma sub_div (h : 0 < b → b < a → c ≠ 0) : (a - b) / c = a / c - b / c := by { simp_rw div_eq_mul_inv, exact ennreal.sub_mul (by simpa using h) } From 2bb0d084deca5cc64bf4dc95d53da50f36b28dd4 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Sat, 7 Jan 2023 11:52:46 +0000 Subject: [PATCH 0171/1291] feat(probability/kernel/basic): measurability of the integral against a kernel (#18061) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The main result is ```lean theorem measurable_lintegral (κ : kernel α β) [is_s_finite_kernel κ] {f : α → β → ℝ≥0∞} (hf : measurable (function.uncurry f)) : measurable (λ a, ∫⁻ b, f a b ∂(κ a)) ``` --- src/probability/kernel/basic.lean | 136 ++++++++++++++++++++++++++++++ 1 file changed, 136 insertions(+) diff --git a/src/probability/kernel/basic.lean b/src/probability/kernel/basic.lean index 48a1852a37ef7..6dba6cfff2e73 100644 --- a/src/probability/kernel/basic.lean +++ b/src/probability/kernel/basic.lean @@ -32,6 +32,10 @@ Classes of kernels: * `ext_fun`: if `∫⁻ b, f b ∂(κ a) = ∫⁻ b, f b ∂(η a)` for all measurable functions `f` and all `a`, then the two kernels `κ` and `η` are equal. +* `measurable_lintegral`: the function `a ↦ ∫⁻ b, f a b ∂(κ a)` is measurable, for an s-finite + kernel `κ : kernel α β` and a function `f : α → β → ℝ≥0∞` such that `function.uncurry f` + is measurable. + -/ open measure_theory @@ -400,6 +404,138 @@ end end restrict +section measurable_lintegral + +/-- This is an auxiliary lemma for `measurable_prod_mk_mem`. -/ +lemma measurable_prod_mk_mem_of_finite (κ : kernel α β) {t : set (α × β)} (ht : measurable_set t) + (hκs : ∀ a, is_finite_measure (κ a)) : + measurable (λ a, κ a {b | (a, b) ∈ t}) := +begin + -- `t` is a measurable set in the product `α × β`: we use that the product σ-algebra is generated + -- by boxes to prove the result by induction. + refine measurable_space.induction_on_inter generate_from_prod.symm is_pi_system_prod _ _ _ _ ht, + { -- case `t = ∅` + simp only [set.mem_empty_iff_false, set.set_of_false, measure_empty, measurable_const], }, + { -- case of a box: `t = t₁ ×ˢ t₂` for measurable sets `t₁` and `t₂` + intros t' ht', + simp only [set.mem_image2, set.mem_set_of_eq, exists_and_distrib_left] at ht', + obtain ⟨t₁, ht₁, t₂, ht₂, rfl⟩ := ht', + simp only [set.prod_mk_mem_set_prod_eq], + classical, + have h_eq_ite : (λ a, κ a {b : β | a ∈ t₁ ∧ b ∈ t₂}) = λ a, ite (a ∈ t₁) (κ a t₂) 0, + { ext1 a, + split_ifs, + { simp only [h, true_and], refl, }, + { simp only [h, false_and, set.set_of_false, set.inter_empty, measure_empty], }, }, + rw h_eq_ite, + exact measurable.ite ht₁ (kernel.measurable_coe κ ht₂) measurable_const }, + { -- we assume that the result is true for `t` and we prove it for `tᶜ` + intros t' ht' h_meas, + have h_eq_sdiff : ∀ a, {b : β | (a, b) ∈ t'ᶜ} = set.univ \ {b : β | (a, b) ∈ t'}, + { intro a, + ext1 b, + simp only [set.mem_compl_iff, set.mem_set_of_eq, set.mem_diff, set.mem_univ, true_and], }, + simp_rw h_eq_sdiff, + have : (λ a, κ a (set.univ \ {b : β | (a, b) ∈ t'})) + = (λ a, (κ a set.univ - κ a {b : β | (a, b) ∈ t'})), + { ext1 a, + rw [← set.diff_inter_self_eq_diff, set.inter_univ, measure_diff], + { exact set.subset_univ _, }, + { exact (@measurable_prod_mk_left α β _ _ a) t' ht', }, + { exact measure_ne_top _ _, }, }, + rw this, + exact measurable.sub (kernel.measurable_coe κ measurable_set.univ) h_meas, }, + { -- we assume that the result is true for a family of disjoint sets and prove it for their union + intros f h_disj hf_meas hf, + have h_Union : (λ a, κ a {b : β | (a, b) ∈ ⋃ i, f i}) = λ a, κ a (⋃ i, {b : β | (a, b) ∈ f i}), + { ext1 a, + congr' with b, + simp only [set.mem_Union, set.supr_eq_Union, set.mem_set_of_eq], + refl, }, + rw h_Union, + have h_tsum : (λ a, κ a (⋃ i, {b : β | (a, b) ∈ f i})) = λ a, ∑' i, κ a {b : β | (a, b) ∈ f i}, + { ext1 a, + rw measure_Union, + { intros i j hij s hsi hsj b hbs, + have habi : {(a, b)} ⊆ f i, by { rw set.singleton_subset_iff, exact hsi hbs, }, + have habj : {(a, b)} ⊆ f j, by { rw set.singleton_subset_iff, exact hsj hbs, }, + simpa only [set.bot_eq_empty, set.le_eq_subset, set.singleton_subset_iff, + set.mem_empty_iff_false] using h_disj hij habi habj, }, + { exact λ i, (@measurable_prod_mk_left α β _ _ a) _ (hf_meas i), }, }, + rw h_tsum, + exact measurable.ennreal_tsum hf, }, +end + +lemma measurable_prod_mk_mem (κ : kernel α β) [is_s_finite_kernel κ] + {t : set (α × β)} (ht : measurable_set t) : + measurable (λ a, κ a {b | (a, b) ∈ t}) := +begin + rw ← kernel_sum_seq κ, + have : ∀ a, kernel.sum (seq κ) a {b : β | (a, b) ∈ t} = ∑' n, seq κ n a {b : β | (a, b) ∈ t}, + from λ a, kernel.sum_apply' _ _ (measurable_prod_mk_left ht), + simp_rw this, + refine measurable.ennreal_tsum (λ n, _), + exact measurable_prod_mk_mem_of_finite (seq κ n) ht infer_instance, +end + +lemma measurable_lintegral_indicator_const (κ : kernel α β) [is_s_finite_kernel κ] + {t : set (α × β)} (ht : measurable_set t) (c : ℝ≥0∞) : + measurable (λ a, ∫⁻ b, t.indicator (function.const (α × β) c) (a, b) ∂(κ a)) := +begin + simp_rw lintegral_indicator_const_comp measurable_prod_mk_left ht _, + exact measurable.const_mul (measurable_prod_mk_mem _ ht) c, +end + +/-- For an s-finite kernel `κ` and a function `f : α → β → ℝ≥0∞` which is measurable when seen as a +map from `α × β` (hypothesis `measurable (function.uncurry f)`), the integral +`a ↦ ∫⁻ b, f a b ∂(κ a)` is measurable. -/ +theorem measurable_lintegral (κ : kernel α β) [is_s_finite_kernel κ] + {f : α → β → ℝ≥0∞} (hf : measurable (function.uncurry f)) : + measurable (λ a, ∫⁻ b, f a b ∂(κ a)) := +begin + let F : ℕ → simple_func (α × β) ℝ≥0∞ := simple_func.eapprox (function.uncurry f), + have h : ∀ a, (⨆ n, F n a) = function.uncurry f a, + from simple_func.supr_eapprox_apply (function.uncurry f) hf, + simp only [prod.forall, function.uncurry_apply_pair] at h, + simp_rw ← h, + have : ∀ a, ∫⁻ b, (⨆ n, F n (a, b)) ∂(κ a) = ⨆ n, ∫⁻ b, F n (a, b) ∂(κ a), + { intro a, + rw lintegral_supr, + { exact λ n, (F n).measurable.comp measurable_prod_mk_left, }, + { exact λ i j hij b, simple_func.monotone_eapprox (function.uncurry f) hij _, }, }, + simp_rw this, + refine measurable_supr (λ n, simple_func.induction _ _ (F n)), + { intros c t ht, + simp only [simple_func.const_zero, simple_func.coe_piecewise, simple_func.coe_const, + simple_func.coe_zero, set.piecewise_eq_indicator], + exact measurable_lintegral_indicator_const κ ht c, }, + { intros g₁ g₂ h_disj hm₁ hm₂, + simp only [simple_func.coe_add, pi.add_apply], + have h_add : (λ a, ∫⁻ b, g₁ (a, b) + g₂ (a, b) ∂(κ a)) + = (λ a, ∫⁻ b, g₁ (a, b) ∂(κ a)) + (λ a, ∫⁻ b, g₂ (a, b) ∂(κ a)), + { ext1 a, + rw [pi.add_apply, lintegral_add_left (g₁.measurable.comp measurable_prod_mk_left)], }, + rw h_add, + exact measurable.add hm₁ hm₂, }, +end + +lemma measurable_lintegral' (κ : kernel α β) [is_s_finite_kernel κ] + {f : β → ℝ≥0∞} (hf : measurable f) : + measurable (λ a, ∫⁻ b, f b ∂(κ a)) := +measurable_lintegral κ (hf.comp measurable_snd) + +lemma measurable_set_lintegral (κ : kernel α β) [is_s_finite_kernel κ] + {f : α → β → ℝ≥0∞} (hf : measurable (function.uncurry f)) {s : set β} (hs : measurable_set s) : + measurable (λ a, ∫⁻ b in s, f a b ∂(κ a)) := +by { simp_rw ← lintegral_restrict κ hs, exact measurable_lintegral _ hf } + +lemma measurable_set_lintegral' (κ : kernel α β) [is_s_finite_kernel κ] + {f : β → ℝ≥0∞} (hf : measurable f) {s : set β} (hs : measurable_set s) : + measurable (λ a, ∫⁻ b in s, f b ∂(κ a)) := +measurable_set_lintegral κ (hf.comp measurable_snd) hs + +end measurable_lintegral + end kernel end probability_theory From 9d81ef7e7004ae91a69bdd86bddc05a744ae8116 Mon Sep 17 00:00:00 2001 From: sgouezel Date: Sat, 7 Jan 2023 13:53:09 +0000 Subject: [PATCH 0172/1291] feat(analysis/convolution): regularity of convolution for functions depending on a parameter (#17626) Show that the convolution of `f` and `g` is smooth when `g` is. A version of this statement is already in mathlib, but in this PR we prove the version where `g` depends on a parameter. Co-authored-by: Floris van Doorn --- src/analysis/convolution.lean | 574 ++++++++++++++++++++++++++++++---- src/topology/support.lean | 2 +- 2 files changed, 514 insertions(+), 62 deletions(-) diff --git a/src/analysis/convolution.lean b/src/analysis/convolution.lean index cd4b826c51faf..44f2aab93c6ff 100644 --- a/src/analysis/convolution.lean +++ b/src/analysis/convolution.lean @@ -60,6 +60,9 @@ This generality has several advantages * `has_compact_support.cont_diff_convolution_right` and `has_compact_support.cont_diff_convolution_left`: the convolution is `𝒞ⁿ` if one of the functions is `𝒞ⁿ` with compact support and the other function in locally integrable. + +Versions of these statements for functions depending on a parameter are also given. + * `convolution_tendsto_right`: Given a sequence of nonnegative normalized functions whose support tends to a small neighborhood around `0`, the convolution tends to the right argument. This is specialized to bump functions in `cont_diff_bump_of_inner.convolution_tendsto_right`. @@ -87,9 +90,14 @@ open set function filter measure_theory measure_theory.measure topological_space open continuous_linear_map metric open_locale pointwise topological_space nnreal filter -variables {𝕜 G E E' E'' F F' F'' : Type*} +universes u𝕜 uG uE uE' uE'' uF uF' uF'' uP + +variables {𝕜 : Type u𝕜} {G : Type uG} {E : Type uE} {E' : Type uE'} {E'' : Type uE''} +{F : Type uF} {F' : Type uF'} {F'' : Type uF''} {P : Type uP} + variables [normed_add_comm_group E] [normed_add_comm_group E'] [normed_add_comm_group E''] - [normed_add_comm_group F] {f f' : G → E} {g g' : G → E'} {x x' : G} {y y' : E} + [normed_add_comm_group F] + {f f' : G → E} {g g' : G → E'} {x x' : G} {y y' : E} section nontrivially_normed_field @@ -101,15 +109,13 @@ section no_measurability variables [add_group G] [topological_space G] -lemma has_compact_support.convolution_integrand_bound_right_of_subset (hcg : has_compact_support g) - (hg : continuous g) {x t : G} {s u : set G} (hx : x ∈ s) (hu : - tsupport g + s ⊆ u) : - ‖L (f t) (g (x - t))‖ ≤ u.indicator (λ t, ‖L‖ * ‖f t‖ * (⨆ i, ‖g i‖)) t := +lemma convolution_integrand_bound_right_of_le_of_subset + {C : ℝ} (hC : ∀ i, ‖g i‖ ≤ C) {x t : G} {s u : set G} (hx : x ∈ s) (hu : - tsupport g + s ⊆ u) : + ‖L (f t) (g (x - t))‖ ≤ u.indicator (λ t, ‖L‖ * ‖f t‖ * C) t := begin refine le_indicator (λ t ht, _) (λ t ht, _) t, { refine (L.le_op_norm₂ _ _).trans _, - exact mul_le_mul_of_nonneg_left - (le_csupr (hg.norm.bdd_above_range_of_has_compact_support hcg.norm) $ x - t) - (mul_nonneg (norm_nonneg _) (norm_nonneg _)) }, + apply mul_le_mul_of_nonneg_left (hC _) (mul_nonneg (norm_nonneg _) (norm_nonneg _)) }, { have : x - t ∉ support g, { refine mt (λ hxt, _) ht, apply hu, @@ -118,6 +124,14 @@ begin rw [nmem_support.mp this, (L _).map_zero, norm_zero] } end +lemma has_compact_support.convolution_integrand_bound_right_of_subset (hcg : has_compact_support g) + (hg : continuous g) {x t : G} {s u : set G} (hx : x ∈ s) (hu : - tsupport g + s ⊆ u) : + ‖L (f t) (g (x - t))‖ ≤ u.indicator (λ t, ‖L‖ * ‖f t‖ * (⨆ i, ‖g i‖)) t := +begin + apply convolution_integrand_bound_right_of_le_of_subset _ (λ i, _) hx hu, + exact le_csupr (hg.norm.bdd_above_range_of_has_compact_support hcg.norm) _, +end + lemma has_compact_support.convolution_integrand_bound_right (hcg : has_compact_support g) (hg : continuous g) {x t : G} {s : set G} (hx : x ∈ s) : ‖L (f t) (g (x - t))‖ ≤ (- tsupport g + s).indicator (λ t, ‖L‖ * ‖f t‖ * (⨆ i, ‖g i‖)) t := @@ -513,35 +527,158 @@ is_compact_of_is_closed_subset (hcg.is_compact.add hcf) is_closed_closure $ clos ((support_convolution_subset_swap L).trans $ add_subset_add subset_closure subset_closure) (hcg.is_compact.add hcf).is_closed -variables [borel_space G] - -/-- The convolution is continuous if one function is locally integrable and the other has compact -support and is continuous. -/ -lemma has_compact_support.continuous_convolution_right [first_countable_topology G] - (hcg : has_compact_support g) (hf : locally_integrable f μ) - (hg : continuous g) : continuous (f ⋆[L, μ] g) := +variables [borel_space G] [first_countable_topology G] +[topological_space P] [first_countable_topology P] + +/-- The convolution `f * g` is continuous if `f` is locally integrable and `g` is continuous and +compactly supported. Version where `g` depends on an additional parameter in a subset `s` of +a parameter space `P` (and the compact support `k` is independent of the parameter in `s`), +not assuming `t2_space G`. -/ +lemma continuous_on_convolution_right_with_param' + {g : P → G → E'} {s : set P} {k : set G} (hk : is_compact k) (h'k : is_closed k) + (hgs : ∀ p, ∀ x, p ∈ s → x ∉ k → g p x = 0) + (hf : locally_integrable f μ) (hg : continuous_on (↿g) (s ×ˢ univ)) : + continuous_on (λ (q : P × G), (f ⋆[L, μ] g q.1) q.2) (s ×ˢ univ) := begin - refine continuous_iff_continuous_at.mpr (λ x₀, _), - let K' := - tsupport g + {x₀}, - have hK' : is_compact K' := hcg.neg.add is_compact_singleton, + assume q₀ hq₀, + replace hq₀ : q₀.1 ∈ s, by simpa only [mem_prod, mem_univ, and_true] using hq₀, + have A : ∀ p ∈ s, continuous (g p), + { assume p hp, + apply hg.comp_continuous (continuous_const.prod_mk continuous_id') (λ x, _), + simpa only [prod_mk_mem_set_prod_eq, mem_univ, and_true] using hp }, + have B : ∀ p ∈ s, tsupport (g p) ⊆ k := + λ p hp, closure_minimal (support_subset_iff'.2 (λ z hz, hgs _ _ hp hz)) h'k, + /- We find a small neighborhood of `{q₀.1} × k` on which the function is uniformly bounded. + This follows from the continuity at all points of the compact set `k`. -/ + obtain ⟨w, C, w_open, q₀w, Cnonneg, hw⟩ : + ∃ w C, is_open w ∧ q₀.1 ∈ w ∧ 0 ≤ C ∧ ∀ p x, p ∈ w ∩ s → ‖g p x‖ ≤ C, + { have A : is_compact ({q₀.1} ×ˢ k), from is_compact_singleton.prod hk, + obtain ⟨t, kt, t_open, ht⟩ : + ∃ t, {q₀.1} ×ˢ k ⊆ t ∧ is_open t ∧ bounded (↿g '' (t ∩ s ×ˢ univ)), + { apply exists_is_open_bounded_image_inter_of_is_compact_of_continuous_on A _ hg, + simp only [prod_subset_prod_iff, hq₀, singleton_subset_iff, subset_univ, and_self, true_or] }, + obtain ⟨C, Cpos, hC⟩ : ∃ C, 0 < C ∧ (↿g) '' (t ∩ s ×ˢ univ) ⊆ closed_ball (0 : E') C, + from ht.subset_ball_lt 0 0, + obtain ⟨w, w_open, q₀w, hw⟩ : ∃ w, is_open w ∧ q₀.1 ∈ w ∧ w ×ˢ k ⊆ t, + { obtain ⟨w, v, w_open, v_open, hw, hv, hvw⟩ : + ∃ (w : set P) (v : set G), is_open w ∧ is_open v ∧ {q₀.fst} ⊆ w ∧ k ⊆ v ∧ w ×ˢ v ⊆ t, + from generalized_tube_lemma is_compact_singleton hk t_open kt, + exact ⟨w, w_open, singleton_subset_iff.1 hw, + subset.trans (set.prod_mono subset.rfl hv) hvw⟩ }, + refine ⟨w, C, w_open, q₀w, Cpos.le, _⟩, + rintros p x ⟨hp, hps⟩, + by_cases hx : x ∈ k, + { have H : (p, x) ∈ t, + { apply hw, + simp only [prod_mk_mem_set_prod_eq, hp, hx, and_true], }, + have H' : (p, x) ∈ (s ×ˢ univ : set (P × G)), + by simpa only [prod_mk_mem_set_prod_eq, mem_univ, and_true] using hps, + have : g p x ∈ closed_ball (0 : E') C, from hC (mem_image_of_mem _ ⟨H, H'⟩), + rwa mem_closed_ball_zero_iff at this }, + { have : g p x = 0, from hgs _ _ hps hx, + rw this, + simpa only [norm_zero] using Cpos.le } }, + have I1 : ∀ᶠ (q : P × G) in 𝓝[s ×ˢ univ] q₀, + ae_strongly_measurable (λ (a : G), L (f a) (g q.1 (q.2 - a))) μ, + { filter_upwards [self_mem_nhds_within], + rintros ⟨p, x⟩ ⟨hp, hx⟩, + refine (has_compact_support.convolution_exists_right L _ hf (A _ hp) _).1, + exact is_compact_of_is_closed_subset hk (is_closed_tsupport _) (B p hp) }, + let K' := - k + {q₀.2}, + have hK' : is_compact K' := hk.neg.add is_compact_singleton, obtain ⟨U, U_open, K'U, hU⟩ : ∃ U, is_open U ∧ K' ⊆ U ∧ integrable_on f U μ, from hf.integrable_on_nhds_is_compact hK', - have : ∀ᶠ x in 𝓝 x₀, ∀ᵐ (t : G) ∂μ, - ‖L (f t) (g (x - t))‖ ≤ U.indicator (λ t, ‖L‖ * ‖f t‖ * (⨆ i, ‖g i‖)) t, + let bound : G → ℝ := indicator U (λ a, ‖L‖ * ‖f a‖ * C), + have I2 : ∀ᶠ (q : P × G) in 𝓝[s ×ˢ univ] q₀, ∀ᵐ a ∂μ, ‖L (f a) (g q.1 (q.2 - a))‖ ≤ bound a, { obtain ⟨V, V_mem, hV⟩ : ∃ (V : set G) (H : V ∈ 𝓝 (0 : G)), K' + V ⊆ U, from compact_open_separated_add_right hK' U_open K'U, - have : {x₀} + V ∈ 𝓝 x₀, from singleton_add_mem_nhds_of_nhds_zero x₀ V_mem, - filter_upwards [this] with x hx, - apply eventually_of_forall (λ t, _), - apply hcg.convolution_integrand_bound_right_of_subset L hg hx, - rwa ← add_assoc }, - refine continuous_at_of_dominated _ this _ _, - { apply eventually_of_forall (λ x, _), - exact (has_compact_support.convolution_exists_right L hcg hf hg x).ae_strongly_measurable }, + have : ((w ∩ s) ×ˢ ({q₀.2} + V) : set (P × G)) ∈ 𝓝[s ×ˢ univ] q₀, + { conv_rhs { rw [← @prod.mk.eta _ _ q₀, nhds_within_prod_eq, nhds_within_univ] }, + refine filter.prod_mem_prod _ (singleton_add_mem_nhds_of_nhds_zero q₀.2 V_mem), + exact mem_nhds_within_iff_exists_mem_nhds_inter.2 ⟨w, w_open.mem_nhds q₀w, subset.rfl⟩ }, + filter_upwards [this], + rintros ⟨p, x⟩ hpx, + simp only [prod_mk_mem_set_prod_eq] at hpx, + apply eventually_of_forall (λ a, _), + apply convolution_integrand_bound_right_of_le_of_subset _ _ hpx.2 _, + { assume x, + exact hw _ _ hpx.1 }, + { rw ← add_assoc, + apply subset.trans (add_subset_add_right (add_subset_add_right _)) hV, + rw neg_subset_neg, + exact B p hpx.1.2 } }, + have I3 : integrable bound μ, { rw [integrable_indicator_iff U_open.measurable_set], exact (hU.norm.const_mul _).mul_const _ }, - { exact eventually_of_forall (λ t, (L.continuous₂.comp₂ continuous_const $ - hg.comp $ continuous_id.sub $ by apply continuous_const).continuous_at) } + have I4 : ∀ᵐ (a : G) ∂μ, continuous_within_at (λ (q : P × G), L (f a) (g q.1 (q.2 - a))) + (s ×ˢ univ) q₀, + { apply eventually_of_forall (λ a, _), + suffices H : continuous_within_at (λ (q : P × G), (f a, g q.1 (q.2 - a))) (s ×ˢ univ) q₀, + from L.continuous₂.continuous_at.comp_continuous_within_at H, + apply continuous_within_at_const.prod, + change continuous_within_at (λ (q : P × G), ↿g (q.1, q.2 - a)) (s ×ˢ univ) q₀, + have : continuous_at (λ (q : P × G), (q.1, q.2 - a)) (q₀.1, q₀.2), + from (continuous_fst.prod_mk (continuous_snd.sub continuous_const)).continuous_at, + rw ← @prod.mk.eta _ _ q₀, + have h'q₀ : (q₀.1, q₀.2 - a) ∈ (s ×ˢ univ : set (P × G)) := ⟨hq₀, mem_univ _⟩, + refine continuous_within_at.comp (hg _ h'q₀) this.continuous_within_at _, + rintros ⟨q, x⟩ ⟨hq, hx⟩, + exact ⟨hq, mem_univ _⟩ }, + exact continuous_within_at_of_dominated I1 I2 I3 I4, +end + +/-- The convolution `f * g` is continuous if `f` is locally integrable and `g` is continuous and +compactly supported. Version where `g` depends on an additional parameter in a subset `s` of +a parameter space `P` (and the compact support `k` is independent of the parameter in `s`). -/ +lemma continuous_on_convolution_right_with_param + [t2_space G] {g : P → G → E'} {s : set P} {k : set G} (hk : is_compact k) + (hgs : ∀ p, ∀ x, p ∈ s → x ∉ k → g p x = 0) + (hf : locally_integrable f μ) (hg : continuous_on (↿g) (s ×ˢ univ)) : + continuous_on (λ (q : P × G), (f ⋆[L, μ] g q.1) q.2) (s ×ˢ univ) := +continuous_on_convolution_right_with_param' L hk hk.is_closed hgs hf hg + +/-- The convolution `f * g` is continuous if `f` is locally integrable and `g` is continuous and +compactly supported. Version where `g` depends on an additional parameter in an open subset `s` of +a parameter space `P` (and the compact support `k` is independent of the parameter in `s`), +given in terms of compositions with an additional continuous map. +Version not assuming `t2_space G`. -/ +lemma continuous_on_convolution_right_with_param_comp' + {s : set P} {v : P → G} (hv : continuous_on v s) + {g : P → G → E'} {k : set G} + (hk : is_compact k) (h'k : is_closed k) (hgs : ∀ p, ∀ x, p ∈ s → x ∉ k → g p x = 0) + (hf : locally_integrable f μ) (hg : continuous_on (↿g) (s ×ˢ univ)) : + continuous_on (λ x, (f ⋆[L, μ] g x) (v x)) s := +begin + apply (continuous_on_convolution_right_with_param' L hk h'k hgs hf hg).comp + (continuous_on_id.prod hv), + assume x hx, + simp only [hx, prod_mk_mem_set_prod_eq, mem_univ, and_self, id.def], +end + +/-- The convolution `f * g` is continuous if `f` is locally integrable and `g` is continuous and +compactly supported. Version where `g` depends on an additional parameter in an open subset `s` of +a parameter space `P` (and the compact support `k` is independent of the parameter in `s`), +given in terms of compositions with an additional continuous map. -/ +lemma continuous_on_convolution_right_with_param_comp [t2_space G] + {s : set P} {v : P → G} (hv : continuous_on v s) + {g : P → G → E'} {k : set G} + (hk : is_compact k) (hgs : ∀ p, ∀ x, p ∈ s → x ∉ k → g p x = 0) + (hf : locally_integrable f μ) (hg : continuous_on (↿g) (s ×ˢ univ)) : + continuous_on (λ x, (f ⋆[L, μ] g x) (v x)) s := +continuous_on_convolution_right_with_param_comp' L hv hk hk.is_closed hgs hf hg + +/-- The convolution is continuous if one function is locally integrable and the other has compact +support and is continuous. -/ +lemma has_compact_support.continuous_convolution_right + (hcg : has_compact_support g) (hf : locally_integrable f μ) + (hg : continuous g) : continuous (f ⋆[L, μ] g) := +begin + rw continuous_iff_continuous_on_univ, + let g' : G → G → E' := λ p q, g q, + have : continuous_on (↿g') (univ ×ˢ univ) := (hg.comp continuous_snd).continuous_on, + exact continuous_on_convolution_right_with_param_comp' L + (continuous_iff_continuous_on_univ.1 continuous_id) hcg (is_closed_tsupport _) + (λ p x hp hx, image_eq_zero_of_nmem_tsupport hx) hf this, end /-- The convolution is continuous if one function is integrable and the other is bounded and @@ -1019,37 +1156,6 @@ begin exact hcf.has_fderiv_at_convolution_right L.flip hg hf x₀, end -lemma has_compact_support.cont_diff_convolution_right - (hcg : has_compact_support g) (hf : locally_integrable f μ) (hg : cont_diff 𝕜 n g) : - cont_diff 𝕜 n (f ⋆[L, μ] g) := -begin - rcases hcg.eq_zero_or_finite_dimensional 𝕜 hg.continuous with rfl|fin_dim, - { simp only [convolution_zero], exact cont_diff_zero_fun, }, - resetI, - induction n using enat.nat_induction with n ih ih generalizing g, - { rw [cont_diff_zero] at hg ⊢, - exact hcg.continuous_convolution_right L hf hg }, - { have h : ∀ x, has_fderiv_at (f ⋆[L, μ] g) ((f ⋆[L.precompR G, μ] fderiv 𝕜 g) x) x := - hcg.has_fderiv_at_convolution_right L hf hg.one_of_succ, - rw cont_diff_succ_iff_fderiv_apply, - split, - { exact λ x₀, ⟨_, h x₀⟩ }, - { simp_rw [fderiv_eq h, convolution_precompR_apply L hf (hcg.fderiv 𝕜) - (hg.one_of_succ.continuous_fderiv le_rfl)], - intro x, - refine ih _ _, - { refine @has_compact_support.comp_left _ _ _ _ _ _ (λ (G : _ →L[𝕜] _), G x) _ - (hcg.fderiv 𝕜) (continuous_linear_map.zero_apply x) }, - { revert x, rw [← cont_diff_clm_apply_iff], - exact (cont_diff_succ_iff_fderiv.mp hg).2 } } }, - { rw [cont_diff_top] at hg ⊢, exact λ n, ih n hcg (hg n) } -end - -lemma has_compact_support.cont_diff_convolution_left [is_neg_invariant μ] - (hcf : has_compact_support f) (hf : cont_diff 𝕜 n f) (hg : locally_integrable g μ) : - cont_diff 𝕜 n (f ⋆[L, μ] g) := -by { rw [← convolution_flip], exact hcf.cont_diff_convolution_right L.flip hg hf } - end is_R_or_C section real @@ -1086,3 +1192,349 @@ begin end end real + +section with_param + +variables [is_R_or_C 𝕜] [normed_space 𝕜 E] [normed_space 𝕜 E'] [normed_space 𝕜 E''] +[normed_space ℝ F] [normed_space 𝕜 F] [complete_space F] [measurable_space G] +[normed_add_comm_group G] [borel_space G] [normed_space 𝕜 G] +[normed_add_comm_group P] [normed_space 𝕜 P] +{μ : measure G} (L : E →L[𝕜] E' →L[𝕜] F) + +/-- The derivative of the convolution `f * g` is given by `f * Dg`, when `f` is locally integrable +and `g` is `C^1` and compactly supported. Version where `g` depends on an additional parameter in an +open subset `s` of a parameter space `P` (and the compact support `k` is independent of the +parameter in `s`). -/ +lemma has_fderiv_at_convolution_right_with_param + {g : P → G → E'} {s : set P} {k : set G} (hs : is_open s) (hk : is_compact k) + (hgs : ∀ p, ∀ x, p ∈ s → x ∉ k → g p x = 0) + (hf : locally_integrable f μ) (hg : cont_diff_on 𝕜 1 ↿g (s ×ˢ univ)) + (q₀ : P × G) (hq₀ : q₀.1 ∈ s) : + has_fderiv_at (λ (q : P × G), (f ⋆[L, μ] g q.1) q.2) + ((f ⋆[L.precompR (P × G), μ] (λ (x : G), fderiv 𝕜 ↿g (q₀.1, x))) q₀.2) q₀ := +begin + let g' := fderiv 𝕜 ↿g, + have A : ∀ p ∈ s, continuous (g p), + { assume p hp, + apply hg.continuous_on.comp_continuous (continuous_const.prod_mk continuous_id') (λ x, _), + simpa only [prod_mk_mem_set_prod_eq, mem_univ, and_true] using hp }, + have A' : ∀ (q : P × G), q.1 ∈ s → s ×ˢ univ ∈ 𝓝 q, + { assume q hq, + apply (hs.prod is_open_univ).mem_nhds, + simpa only [mem_prod, mem_univ, and_true] using hq }, + /- The derivative of `g` vanishes away from `k`. -/ + have g'_zero : ∀ p x, p ∈ s → x ∉ k → g' (p, x) = 0, + { assume p x hp hx, + refine (has_fderiv_at_zero_of_eventually_const 0 _).fderiv, + have M2 : kᶜ ∈ 𝓝 x, from is_open.mem_nhds hk.is_closed.is_open_compl hx, + have M1 : s ∈ 𝓝 p, from hs.mem_nhds hp, + rw nhds_prod_eq, + filter_upwards [prod_mem_prod M1 M2], + rintros ⟨p, y⟩ ⟨hp, hy⟩, + exact hgs p y hp hy }, + /- We find a small neighborhood of `{q₀.1} × k` on which the derivative is uniformly bounded. This + follows from the continuity at all points of the compact set `k`. -/ + obtain ⟨ε, C, εpos, Cnonneg, h₀ε, hε⟩ : + ∃ ε C, 0 < ε ∧ 0 ≤ C ∧ ball q₀.1 ε ⊆ s ∧ ∀ p x, ‖p - q₀.1‖ < ε → ‖g' (p, x)‖ ≤ C, + { have A : is_compact ({q₀.1} ×ˢ k), from is_compact_singleton.prod hk, + obtain ⟨t, kt, t_open, ht⟩ : ∃ t, {q₀.1} ×ˢ k ⊆ t ∧ is_open t ∧ bounded (g' '' t), + { have B : continuous_on g' (s ×ˢ univ), + from hg.continuous_on_fderiv_of_open (hs.prod is_open_univ) le_rfl, + apply exists_is_open_bounded_image_of_is_compact_of_continuous_on A + (hs.prod is_open_univ) _ B, + simp only [prod_subset_prod_iff, hq₀, singleton_subset_iff, subset_univ, and_self, true_or] }, + obtain ⟨ε, εpos, hε, h'ε⟩ : + ∃ (ε : ℝ), 0 < ε ∧ thickening ε ({q₀.fst} ×ˢ k) ⊆ t ∧ ball q₀.1 ε ⊆ s, + { obtain ⟨ε, εpos, hε⟩ : ∃ (ε : ℝ), 0 < ε ∧ thickening ε ({q₀.fst} ×ˢ k) ⊆ t, + from A.exists_thickening_subset_open t_open kt, + obtain ⟨δ, δpos, hδ⟩ : ∃ (δ : ℝ) (H : 0 < δ), ball q₀.1 δ ⊆ s, + from metric.is_open_iff.1 hs _ hq₀, + refine ⟨min ε δ, lt_min εpos δpos, _, _⟩, + { exact subset.trans (thickening_mono (min_le_left _ _) _) hε }, + { exact subset.trans (ball_subset_ball (min_le_right _ _)) hδ } }, + obtain ⟨C, Cpos, hC⟩ : ∃ C, 0 < C ∧ g' '' t ⊆ closed_ball 0 C, from ht.subset_ball_lt 0 0, + refine ⟨ε, C, εpos, Cpos.le, h'ε, λ p x hp, _⟩, + have hps : p ∈ s, from h'ε (mem_ball_iff_norm.2 hp), + by_cases hx : x ∈ k, + { have H : (p, x) ∈ t, + { apply hε, + refine mem_thickening_iff.2 ⟨(q₀.1, x), _, _⟩, + { simp only [hx, singleton_prod, mem_image, prod.mk.inj_iff, eq_self_iff_true, true_and, + exists_eq_right] }, + { rw ← dist_eq_norm at hp, + simpa only [prod.dist_eq, εpos, dist_self, max_lt_iff, and_true] using hp } }, + have : g' (p, x) ∈ closed_ball (0 : P × G →L[𝕜] E') C, from hC (mem_image_of_mem _ H), + rwa mem_closed_ball_zero_iff at this }, + { have : g' (p, x) = 0, from g'_zero _ _ hps hx, + rw this, + simpa only [norm_zero] using Cpos.le } }, + /- Now, we wish to apply a theorem on differentiation of integrals. For this, we need to check + trivial measurability or integrability assumptions (in `I1`, `I2`, `I3`), as well as a uniform + integrability assumption over the derivative (in `I4` and `I5`) and pointwise differentiability + in `I6`. -/ + have I1 : ∀ᶠ (x : P × G) in 𝓝 q₀, + ae_strongly_measurable (λ (a : G), L (f a) (g x.1 (x.2 - a))) μ, + { filter_upwards [A' q₀ hq₀], + rintros ⟨p, x⟩ ⟨hp, hx⟩, + refine (has_compact_support.convolution_exists_right L _ hf (A _ hp) _).1, + apply is_compact_of_is_closed_subset hk (is_closed_tsupport _), + exact closure_minimal (support_subset_iff'.2 (λ z hz, hgs _ _ hp hz)) hk.is_closed, }, + have I2 : integrable (λ (a : G), L (f a) (g q₀.1 (q₀.2 - a))) μ, + { have M : has_compact_support (g q₀.1), + from has_compact_support.intro hk (λ x hx, hgs q₀.1 x hq₀ hx), + apply M.convolution_exists_right L hf (A q₀.1 hq₀) q₀.2 }, + have I3 : ae_strongly_measurable (λ (a : G), (L (f a)).comp (g' (q₀.fst, q₀.snd - a))) μ, + { have T : has_compact_support (λ y, g' (q₀.1, y)), + from has_compact_support.intro hk (λ x hx, g'_zero q₀.1 x hq₀ hx), + apply (has_compact_support.convolution_exists_right (L.precompR (P × G) : _) T hf _ q₀.2).1, + have : continuous_on g' (s ×ˢ univ), + from hg.continuous_on_fderiv_of_open (hs.prod is_open_univ) le_rfl, + apply this.comp_continuous (continuous_const.prod_mk continuous_id'), + assume x, + simpa only [prod_mk_mem_set_prod_eq, mem_univ, and_true] using hq₀ }, + set K' := - k + {q₀.2} with K'_def, + have hK' : is_compact K' := hk.neg.add is_compact_singleton, + obtain ⟨U, U_open, K'U, hU⟩ : ∃ U, is_open U ∧ K' ⊆ U ∧ integrable_on f U μ, + from hf.integrable_on_nhds_is_compact hK', + obtain ⟨δ, δpos, δε, hδ⟩ : ∃ δ, (0 : ℝ) < δ ∧ δ ≤ ε ∧ K' + ball 0 δ ⊆ U, + { obtain ⟨V, V_mem, hV⟩ : ∃ (V : set G) (V_mem : V ∈ 𝓝 (0 : G)), K' + V ⊆ U, + from compact_open_separated_add_right hK' U_open K'U, + rcases metric.mem_nhds_iff.1 V_mem with ⟨δ, δpos, hδ⟩, + refine ⟨min δ ε, lt_min δpos εpos, min_le_right _ _, _⟩, + exact (add_subset_add_left ((ball_subset_ball (min_le_left _ _)).trans hδ)).trans hV }, + let bound : G → ℝ := indicator U (λ a, ‖L.precompR (P × G)‖ * ‖f a‖ * C), + have I4 : ∀ᵐ (a : G) ∂μ, ∀ (x : P × G), dist x q₀ < δ → + ‖L.precompR (P × G) (f a) (g' (x.fst, x.snd - a))‖ ≤ bound a, + { apply eventually_of_forall, + assume a x hx, + rw [prod.dist_eq, dist_eq_norm, dist_eq_norm] at hx, + have : -tsupport (λ a, g' (x.1, a)) + ball q₀.2 δ ⊆ U, + { apply subset.trans _ hδ, + rw [K'_def, add_assoc], + apply add_subset_add, + { rw neg_subset_neg, + apply closure_minimal (support_subset_iff'.2 (λ z hz, _)) hk.is_closed, + apply g'_zero x.1 z (h₀ε _) hz, + rw mem_ball_iff_norm, + exact ((le_max_left _ _).trans_lt hx).trans_le δε }, + { simp only [add_ball, thickening_singleton, zero_vadd] } }, + apply convolution_integrand_bound_right_of_le_of_subset _ _ _ this, + { assume y, + exact hε _ _ (((le_max_left _ _).trans_lt hx).trans_le δε) }, + { rw mem_ball_iff_norm, + exact (le_max_right _ _).trans_lt hx } }, + have I5 : integrable bound μ, + { rw [integrable_indicator_iff U_open.measurable_set], + exact (hU.norm.const_mul _).mul_const _ }, + have I6 : ∀ᵐ (a : G) ∂μ, ∀ (x : P × G), dist x q₀ < δ → + has_fderiv_at (λ (x : P × G), L (f a) (g x.1 (x.2 - a))) + ((L (f a)).comp (g' (x.fst, x.snd - a))) x, + { apply eventually_of_forall, + assume a x hx, + apply (L _).has_fderiv_at.comp x, + have N : s ×ˢ univ ∈ 𝓝 (x.1, x.2 - a), + { apply A', + apply h₀ε, + rw prod.dist_eq at hx, + exact lt_of_lt_of_le (lt_of_le_of_lt (le_max_left _ _) hx) δε }, + have Z := ((hg.differentiable_on le_rfl).differentiable_at N).has_fderiv_at, + have Z' : has_fderiv_at (λ (x : P × G), (x.1, x.2 - a)) (continuous_linear_map.id 𝕜 (P × G)) x, + { have : (λ (x : P × G), (x.1, x.2 - a)) = id - (λ x, (0, a)), + { ext x; simp only [pi.sub_apply, id.def, prod.fst_sub, sub_zero, prod.snd_sub] }, + simp_rw [this], + exact (has_fderiv_at_id x).sub_const (0, a) }, + exact Z.comp x Z' }, + exact has_fderiv_at_integral_of_dominated_of_fderiv_le δpos I1 I2 I3 I4 I5 I6, +end + +/-- The convolution `f * g` is `C^n` when `f` is locally integrable and `g` is `C^n` and compactly +supported. Version where `g` depends on an additional parameter in an open subset `s` of a +parameter space `P` (and the compact support `k` is independent of the parameter in `s`). +In this version, all the types belong to the same universe (to get an induction working in the +proof). Use instead `cont_diff_on_convolution_right_with_param`, which removes this restriction. -/ +lemma cont_diff_on_convolution_right_with_param_aux + {G : Type uP} {E' : Type uP} {F : Type uP} {P : Type uP} + [normed_add_comm_group E'] [normed_add_comm_group F] + [normed_space 𝕜 E'] [normed_space ℝ F] [normed_space 𝕜 F] [complete_space F] + [measurable_space G] {μ : measure G} [normed_add_comm_group G] [borel_space G] [normed_space 𝕜 G] + [normed_add_comm_group P] [normed_space 𝕜 P] + {f : G → E} {n : ℕ∞} (L : E →L[𝕜] E' →L[𝕜] F) + {g : P → G → E'} + {s : set P} {k : set G} (hs : is_open s) (hk : is_compact k) + (hgs : ∀ p, ∀ x, p ∈ s → x ∉ k → g p x = 0) + (hf : locally_integrable f μ) (hg : cont_diff_on 𝕜 n ↿g (s ×ˢ univ)) : + cont_diff_on 𝕜 n (λ (q : P × G), (f ⋆[L, μ] g q.1) q.2) (s ×ˢ univ) := +begin + /- We have a formula for the derivation of `f * g`, which is of the same form, thanks to + `has_fderiv_at_convolution_right_with_param`. Therefore, we can prove the result by induction on + `n` (but for this we need the spaces at the different steps of the induction to live in the same + universe, which is why we make the assumption in the lemma that all the relevant spaces + come from the same universe). -/ + unfreezingI { induction n using enat.nat_induction with n ih ih generalizing g E' F }, + { rw [cont_diff_on_zero] at hg ⊢, + exact continuous_on_convolution_right_with_param L hk hgs hf hg }, + { let f' : P → G → (P × G →L[𝕜] F) := λ p a, + (f ⋆[L.precompR (P × G), μ] (λ (x : G), fderiv 𝕜 (uncurry g) (p, x))) a, + have A : ∀ (q₀ : P × G), q₀.1 ∈ s → + has_fderiv_at (λ (q : P × G), (f ⋆[L, μ] g q.1) q.2) (f' q₀.1 q₀.2) q₀, + from has_fderiv_at_convolution_right_with_param L hs hk hgs hf hg.one_of_succ, + rw cont_diff_on_succ_iff_fderiv_of_open (hs.prod (@is_open_univ G _)) at ⊢ hg, + split, + { rintros ⟨p, x⟩ ⟨hp, hx⟩, + exact (A (p, x) hp).differentiable_at.differentiable_within_at, }, + { suffices H : cont_diff_on 𝕜 n ↿f' (s ×ˢ univ), + { apply H.congr, + rintros ⟨p, x⟩ ⟨hp, hx⟩, + exact (A (p, x) hp).fderiv }, + have B : ∀ (p : P) (x : G), p ∈ s → x ∉ k → fderiv 𝕜 (uncurry g) (p, x) = 0, + { assume p x hp hx, + apply (has_fderiv_at_zero_of_eventually_const (0 : E') _).fderiv, + have M2 : kᶜ ∈ 𝓝 x, from is_open.mem_nhds hk.is_closed.is_open_compl hx, + have M1 : s ∈ 𝓝 p, from hs.mem_nhds hp, + rw nhds_prod_eq, + filter_upwards [prod_mem_prod M1 M2], + rintros ⟨p, y⟩ ⟨hp, hy⟩, + exact hgs p y hp hy }, + apply ih (L.precompR (P × G) : _) B, + convert hg.2, + apply funext, + rintros ⟨p, x⟩, + refl } }, + { rw [cont_diff_on_top] at hg ⊢, + assume n, + exact ih n L hgs (hg n) } +end + +/-- The convolution `f * g` is `C^n` when `f` is locally integrable and `g` is `C^n` and compactly +supported. Version where `g` depends on an additional parameter in an open subset `s` of a +parameter space `P` (and the compact support `k` is independent of the parameter in `s`). -/ +lemma cont_diff_on_convolution_right_with_param + {f : G → E} {n : ℕ∞} (L : E →L[𝕜] E' →L[𝕜] F) {g : P → G → E'} + {s : set P} {k : set G} (hs : is_open s) (hk : is_compact k) + (hgs : ∀ p, ∀ x, p ∈ s → x ∉ k → g p x = 0) + (hf : locally_integrable f μ) (hg : cont_diff_on 𝕜 n ↿g (s ×ˢ univ)) : + cont_diff_on 𝕜 n (λ (q : P × G), (f ⋆[L, μ] g q.1) q.2) (s ×ˢ univ) := +begin + /- The result is known when all the universes are the same, from + `cont_diff_on_convolution_right_with_param_aux`. We reduce to this situation by pushing + everything through `ulift` continuous linear equivalences. -/ + let eG : Type (max uG uE' uF uP) := ulift G, + borelize eG, + let eE' : Type (max uE' uG uF uP) := ulift E', + let eF : Type (max uF uG uE' uP) := ulift F, + let eP : Type (max uP uG uE' uF) := ulift P, + have isoG : eG ≃L[𝕜] G := continuous_linear_equiv.ulift, + have isoE' : eE' ≃L[𝕜] E' := continuous_linear_equiv.ulift, + have isoF : eF ≃L[𝕜] F := continuous_linear_equiv.ulift, + have isoP : eP ≃L[𝕜] P := continuous_linear_equiv.ulift, + let ef := f ∘ isoG, + let eμ : measure eG := measure.map isoG.symm μ, + let eg : eP → eG → eE' := λ ep ex, isoE'.symm (g (isoP ep) (isoG ex)), + let eL := continuous_linear_map.comp + ((continuous_linear_equiv.arrow_congr isoE' isoF).symm : (E' →L[𝕜] F) →L[𝕜] eE' →L[𝕜] eF) L, + let R := (λ (q : eP × eG), (ef ⋆[eL, eμ] eg q.1) q.2), + have R_contdiff : cont_diff_on 𝕜 n R ((isoP ⁻¹' s) ×ˢ univ), + { have hek : is_compact (isoG ⁻¹' k), + from isoG.to_homeomorph.closed_embedding.is_compact_preimage hk, + have hes : is_open (isoP ⁻¹' s), from isoP.continuous.is_open_preimage _ hs, + refine cont_diff_on_convolution_right_with_param_aux eL hes hek _ _ _, + { assume p x hp hx, + simp only [comp_app, continuous_linear_equiv.prod_apply, linear_isometry_equiv.coe_coe, + continuous_linear_equiv.map_eq_zero_iff], + exact hgs _ _ hp hx }, + { apply (locally_integrable_map_homeomorph isoG.symm.to_homeomorph).2, + convert hf, + ext1 x, + simp only [ef, continuous_linear_equiv.coe_to_homeomorph, comp_app, + continuous_linear_equiv.apply_symm_apply], }, + { apply isoE'.symm.cont_diff.comp_cont_diff_on, + apply hg.comp ((isoP.prod isoG).cont_diff).cont_diff_on, + rintros ⟨p, x⟩ ⟨hp, hx⟩, + simpa only [mem_preimage, continuous_linear_equiv.prod_apply, prod_mk_mem_set_prod_eq, + mem_univ, and_true] using hp } }, + have A : cont_diff_on 𝕜 n (isoF ∘ R ∘ (isoP.prod isoG).symm) (s ×ˢ univ), + { apply isoF.cont_diff.comp_cont_diff_on, + apply R_contdiff.comp (continuous_linear_equiv.cont_diff _).cont_diff_on, + rintros ⟨p, x⟩ ⟨hp, hx⟩, + simpa only [mem_preimage, mem_prod, mem_univ, and_true, continuous_linear_equiv.prod_symm, + continuous_linear_equiv.prod_apply, continuous_linear_equiv.apply_symm_apply] using hp }, + have : isoF ∘ R ∘ (isoP.prod isoG).symm = (λ (q : P × G), (f ⋆[L, μ] g q.1) q.2), + { apply funext, + rintros ⟨p, x⟩, + simp only [R, linear_isometry_equiv.coe_coe, comp_app, continuous_linear_equiv.prod_symm, + continuous_linear_equiv.prod_apply], + simp only [convolution, eL, coe_comp', continuous_linear_equiv.coe_coe, comp_app, eμ], + rw [closed_embedding.integral_map, ← isoF.integral_comp_comm], + swap, { exact isoG.symm.to_homeomorph.closed_embedding }, + congr' 1, + ext1 a, + simp only [ef, eg, comp_app, continuous_linear_equiv.apply_symm_apply, coe_comp', + continuous_linear_equiv.prod_apply, continuous_linear_equiv.map_sub, + continuous_linear_equiv.arrow_congr, continuous_linear_equiv.arrow_congrSL_symm_apply, + continuous_linear_equiv.coe_coe, comp_app, continuous_linear_equiv.apply_symm_apply] }, + simp_rw [this] at A, + exact A, +end + +/-- The convolution `f * g` is `C^n` when `f` is locally integrable and `g` is `C^n` and compactly +supported. Version where `g` depends on an additional parameter in an open subset `s` of a +parameter space `P` (and the compact support `k` is independent of the parameter in `s`), +given in terms of composition with an additional smooth function. -/ +lemma cont_diff_on_convolution_right_with_param_comp + {n : ℕ∞} (L : E →L[𝕜] E' →L[𝕜] F) + {s : set P} {v : P → G} (hv : cont_diff_on 𝕜 n v s) + {f : G → E} {g : P → G → E'} {k : set G} (hs : is_open s) (hk : is_compact k) + (hgs : ∀ p, ∀ x, p ∈ s → x ∉ k → g p x = 0) + (hf : locally_integrable f μ) (hg : cont_diff_on 𝕜 n ↿g (s ×ˢ univ)) : + cont_diff_on 𝕜 n (λ x, (f ⋆[L, μ] g x) (v x)) s := +begin + apply (cont_diff_on_convolution_right_with_param L hs hk hgs hf hg).comp + (cont_diff_on_id.prod hv), + assume x hx, + simp only [hx, mem_preimage, prod_mk_mem_set_prod_eq, mem_univ, and_self, id.def], +end + +/-- The convolution `g * f` is `C^n` when `f` is locally integrable and `g` is `C^n` and compactly +supported. Version where `g` depends on an additional parameter in an open subset `s` of a +parameter space `P` (and the compact support `k` is independent of the parameter in `s`). -/ +lemma cont_diff_on_convolution_left_with_param [μ.is_add_left_invariant] [μ.is_neg_invariant] + (L : E' →L[𝕜] E →L[𝕜] F) {f : G → E} {n : ℕ∞} + {g : P → G → E'} {s : set P} {k : set G} (hs : is_open s) (hk : is_compact k) + (hgs : ∀ p, ∀ x, p ∈ s → x ∉ k → g p x = 0) + (hf : locally_integrable f μ) (hg : cont_diff_on 𝕜 n ↿g (s ×ˢ univ)) : + cont_diff_on 𝕜 n (λ (q : P × G), (g q.1 ⋆[L, μ] f) q.2) (s ×ˢ univ) := +by simpa only [convolution_flip] + using cont_diff_on_convolution_right_with_param L.flip hs hk hgs hf hg + +/-- The convolution `g * f` is `C^n` when `f` is locally integrable and `g` is `C^n` and compactly +supported. Version where `g` depends on an additional parameter in an open subset `s` of a +parameter space `P` (and the compact support `k` is independent of the parameter in `s`), +given in terms of composition with additional smooth functions. -/ +lemma cont_diff_on_convolution_left_with_param_comp [μ.is_add_left_invariant] [μ.is_neg_invariant] + (L : E' →L[𝕜] E →L[𝕜] F) {s : set P} {n : ℕ∞} {v : P → G} (hv : cont_diff_on 𝕜 n v s) + {f : G → E} {g : P → G → E'} {k : set G} (hs : is_open s) (hk : is_compact k) + (hgs : ∀ p, ∀ x, p ∈ s → x ∉ k → g p x = 0) + (hf : locally_integrable f μ) (hg : cont_diff_on 𝕜 n ↿g (s ×ˢ univ)) : + cont_diff_on 𝕜 n (λ x, (g x ⋆[L, μ] f) (v x)) s := +begin + apply (cont_diff_on_convolution_left_with_param L hs hk hgs hf hg).comp (cont_diff_on_id.prod hv), + assume x hx, + simp only [hx, mem_preimage, prod_mk_mem_set_prod_eq, mem_univ, and_self, id.def], +end + +lemma has_compact_support.cont_diff_convolution_right {n : ℕ∞} + (hcg : has_compact_support g) (hf : locally_integrable f μ) (hg : cont_diff 𝕜 n g) : + cont_diff 𝕜 n (f ⋆[L, μ] g) := +begin + rcases exists_compact_iff_has_compact_support.2 hcg with ⟨k, hk, h'k⟩, + rw ← cont_diff_on_univ, + exact cont_diff_on_convolution_right_with_param_comp L cont_diff_on_id is_open_univ hk + (λ p x hp hx, h'k x hx) hf (hg.comp cont_diff_snd).cont_diff_on, +end + +lemma has_compact_support.cont_diff_convolution_left [μ.is_add_left_invariant] [μ.is_neg_invariant] + {n : ℕ∞} (hcf : has_compact_support f) (hf : cont_diff 𝕜 n f) (hg : locally_integrable g μ) : + cont_diff 𝕜 n (f ⋆[L, μ] g) := +by { rw [← convolution_flip], exact hcf.cont_diff_convolution_right L.flip hg hf } + +end with_param diff --git a/src/topology/support.lean b/src/topology/support.lean index 35b0b9e55261e..4beaa91af3bba 100644 --- a/src/topology/support.lean +++ b/src/topology/support.lean @@ -56,7 +56,7 @@ lemma mul_tsupport_eq_empty_iff {f : X → α} : mul_tsupport f = ∅ ↔ f = 1 by rw [mul_tsupport, closure_empty_iff, mul_support_eq_empty_iff] @[to_additive] -lemma image_eq_zero_of_nmem_mul_tsupport {f : X → α} {x : X} (hx : x ∉ mul_tsupport f) : f x = 1 := +lemma image_eq_one_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] From ae0f31c78da8eec342b7f92c35ad02931a676165 Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Sat, 7 Jan 2023 17:16:03 +0000 Subject: [PATCH 0173/1291] feat(ring_theory/ideal/basic): allow noncommutativity in `span_singleton_mul_left_unit` (#18088) --- src/ring_theory/ideal/basic.lean | 17 ++++++++--------- 1 file changed, 8 insertions(+), 9 deletions(-) diff --git a/src/ring_theory/ideal/basic.lean b/src/ring_theory/ideal/basic.lean index 3eb4be6f0504a..4039a1588777d 100644 --- a/src/ring_theory/ideal/basic.lean +++ b/src/ring_theory/ideal/basic.lean @@ -121,6 +121,13 @@ lemma mem_span_singleton' {x y : α} : lemma span_singleton_le_iff_mem {x : α} : span {x} ≤ I ↔ x ∈ I := submodule.span_singleton_le_iff_mem _ _ +lemma span_singleton_mul_left_unit {a : α} (h2 : is_unit a) (x : α) : + span ({a * x} : set α) = span {x} := +begin + apply le_antisymm; rw [span_singleton_le_iff_mem, mem_span_singleton'], + exacts [⟨a, rfl⟩, ⟨_, h2.unit.inv_mul_cancel_left x⟩], +end + lemma span_insert (x) (s : set α) : span (insert x s) = span ({x} : set α) ⊔ span s := submodule.span_insert x s @@ -384,15 +391,7 @@ begin end lemma span_singleton_mul_right_unit {a : α} (h2 : is_unit a) (x : α) : - span ({x * a} : set α) = span {x} := -begin - apply le_antisymm, - { rw span_singleton_le_span_singleton, use a}, - { rw span_singleton_le_span_singleton, rw is_unit.mul_right_dvd h2} -end - -lemma span_singleton_mul_left_unit {a : α} (h2 : is_unit a) (x : α) : - span ({a * x} : set α) = span {x} := by rw [mul_comm, span_singleton_mul_right_unit h2] + span ({x * a} : set α) = span {x} := by rw [mul_comm, span_singleton_mul_left_unit h2] lemma span_singleton_eq_top {x} : span ({x} : set α) = ⊤ ↔ is_unit x := by rw [is_unit_iff_dvd_one, ← span_singleton_le_span_singleton, span_singleton_one, From 940d371319c6658e526349d2c3e1daeeabfae0fd Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Sat, 7 Jan 2023 19:53:52 +0000 Subject: [PATCH 0174/1291] =?UTF-8?q?feat(topology/instances/add=5Fcircle)?= =?UTF-8?q?:=20generalize=20`homeo=5FIcc=5Fquot`=20from=20=E2=84=9D=20to?= =?UTF-8?q?=20archimedean=20groups=20(#17983)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit + Add left/right/two-sided continuity results about `to_Ico/Ioc_mod` and `equiv_Ico`. + Add `tfae_to_Ico_eq_to_Ioc` which states 8 other equivalent conditions for `to_Ico_mod` and `to_Ioc_mod` to agree at a point, from [this comment](https://github.com/leanprover-community/mathlib/pull/17933/files#r1051518394) in #17933. Co-authored-by: David Loeffler Co-authored-by: Junyan Xu Co-authored-by: loefflerd --- src/algebra/order/to_interval_mod.lean | 86 +++++++++++- src/topology/instances/add_circle.lean | 183 +++++++++++++++++++------ 2 files changed, 228 insertions(+), 41 deletions(-) diff --git a/src/algebra/order/to_interval_mod.lean b/src/algebra/order/to_interval_mod.lean index 2f0e79874c36a..bc775f63cb434 100644 --- a/src/algebra/order/to_interval_mod.lean +++ b/src/algebra/order/to_interval_mod.lean @@ -30,7 +30,8 @@ noncomputable theory section linear_ordered_add_comm_group -variables {α : Type*} [linear_ordered_add_comm_group α] [archimedean α] +variables {α : Type*} [linear_ordered_add_comm_group α] [hα : archimedean α] +include hα /-- The unique integer such that this multiple of `b`, added to `x`, is in `Ico a (a + b)`. -/ def to_Ico_div (a : α) {b : α} (hb : 0 < b) (x : α) : ℤ := @@ -457,6 +458,89 @@ begin rw [hz, to_Ioc_mod_zsmul_add] } end +section Ico_Ioc + +variables (a : α) {b : α} (hb : 0 < b) (x : α) + +omit hα +/-- `mem_Ioo_mod a b x` means that `x` lies in the open interval `(a, a + b)` modulo `b`. +Equivalently (as shown below), `x` is not congruent to `a` modulo `b`, or `to_Ico_mod a hb` agrees +with `to_Ioc_mod a hb` at `x`, or `to_Ico_div a hb` agrees with `to_Ioc_div a hb` at `x`. -/ +def mem_Ioo_mod (b x : α) : Prop := ∃ z : ℤ, x + z • b ∈ set.Ioo a (a + b) +include hα + +lemma tfae_mem_Ioo_mod : + tfae [mem_Ioo_mod a b x, + to_Ico_mod a hb x = to_Ioc_mod a hb x, + to_Ico_mod a hb x + b ≠ to_Ioc_mod a hb x, + to_Ico_mod a hb x ≠ a] := +begin + tfae_have : 1 → 2, + { exact λ ⟨i, hi⟩, ((to_Ico_mod_eq_iff hb).2 ⟨hi.1.le, hi.2, i, add_sub_cancel' x _⟩).trans + ((to_Ioc_mod_eq_iff hb).2 ⟨hi.1, hi.2.le, i, add_sub_cancel' x _⟩).symm }, + tfae_have : 2 → 3, + { intro h, rw [h, ne, add_right_eq_self], exact hb.ne' }, + tfae_have : 3 → 4, + { refine mt (λ h, _), rw [h, eq_comm, to_Ioc_mod_eq_iff], + refine ⟨lt_add_of_pos_right a hb, le_rfl, to_Ico_div a hb x + 1, _⟩, + conv_lhs { rw [← h, to_Ico_mod, add_assoc, ← add_one_zsmul, add_sub_cancel'] } }, + tfae_have : 4 → 1, + { have h' := to_Ico_mod_mem_Ico a hb x, exact λ h, ⟨_, h'.1.lt_of_ne' h, h'.2⟩ }, + tfae_finish, +end + +variables {a x} + +lemma mem_Ioo_mod_iff_to_Ico_mod_eq_to_Ioc_mod : + mem_Ioo_mod a b x ↔ to_Ico_mod a hb x = to_Ioc_mod a hb x := (tfae_mem_Ioo_mod a hb x).out 0 1 +lemma mem_Ioo_mod_iff_to_Ico_mod_add_period_ne_to_Ioc_mod : + mem_Ioo_mod a b x ↔ to_Ico_mod a hb x + b ≠ to_Ioc_mod a hb x := (tfae_mem_Ioo_mod a hb x).out 0 2 +lemma mem_Ioo_mod_iff_to_Ico_mod_ne_left : + mem_Ioo_mod a b x ↔ to_Ico_mod a hb x ≠ a := (tfae_mem_Ioo_mod a hb x).out 0 3 +lemma mem_Ioo_mod_iff_to_Ioc_mod_ne_right : mem_Ioo_mod a b x ↔ to_Ioc_mod a hb x ≠ a + b := +begin + rw [mem_Ioo_mod_iff_to_Ico_mod_eq_to_Ioc_mod, to_Ico_mod_eq_iff hb], + obtain ⟨h₁, h₂⟩ := to_Ioc_mod_mem_Ioc a hb x, + exact ⟨λ h, h.2.1.ne, λ h, ⟨h₁.le, h₂.lt_of_ne h, _, add_sub_cancel' x _⟩⟩, +end + +lemma mem_Ioo_mod_iff_to_Ico_div_eq_to_Ioc_div : + mem_Ioo_mod a b x ↔ to_Ico_div a hb x = to_Ioc_div a hb x := +by rw [mem_Ioo_mod_iff_to_Ico_mod_eq_to_Ioc_mod hb, + to_Ico_mod, to_Ioc_mod, add_right_inj, (zsmul_strict_mono_left hb).injective.eq_iff] + +lemma mem_Ioo_mod_iff_to_Ico_div_add_one_ne_to_Ioc_div : + mem_Ioo_mod a b x ↔ to_Ico_div a hb x + 1 ≠ to_Ioc_div a hb x := +by rw [mem_Ioo_mod_iff_to_Ico_mod_add_period_ne_to_Ioc_mod hb, ne, ne, to_Ico_mod, to_Ioc_mod, + add_assoc, add_right_inj, ← add_one_zsmul, (zsmul_strict_mono_left hb).injective.eq_iff] + +include hb + +lemma mem_Ioo_mod_iff_sub_ne_zsmul : mem_Ioo_mod a b x ↔ ∀ z : ℤ, a - x ≠ z • b := +begin + rw [mem_Ioo_mod_iff_to_Ico_mod_ne_left hb, ← not_iff_not], + push_neg, split; intro h, + { rw ← h, exact ⟨_, add_sub_cancel' x _⟩ }, + { exact (to_Ico_mod_eq_iff hb).2 ⟨le_rfl, lt_add_of_pos_right a hb, h⟩ }, +end + +lemma mem_Ioo_mod_iff_eq_mod_zmultiples : + mem_Ioo_mod a b x ↔ (x : α ⧸ add_subgroup.zmultiples b) ≠ a := +begin + rw [mem_Ioo_mod_iff_sub_ne_zsmul hb, ne, eq_comm, + quotient_add_group.eq_iff_sub_mem, add_subgroup.mem_zmultiples_iff], + push_neg, simp_rw ne_comm, +end + +lemma Ico_eq_locus_Ioc_eq_Union_Ioo : + {x | to_Ico_mod a hb x = to_Ioc_mod a hb x} = ⋃ z : ℤ, set.Ioo (a - z • b) (a + b - z • b) := +begin + ext1, simp_rw [set.mem_set_of, set.mem_Union, ← set.add_mem_Ioo_iff_left], + exact (mem_Ioo_mod_iff_to_Ico_mod_eq_to_Ioc_mod hb).symm, +end + +end Ico_Ioc + lemma to_Ico_mod_eq_self {a b x : α} (hb : 0 < b) : to_Ico_mod a hb x = x ↔ a ≤ x ∧ x < a + b := begin rw to_Ico_mod_eq_iff, diff --git a/src/topology/instances/add_circle.lean b/src/topology/instances/add_circle.lean index c9b6076421899..7049ede5b7f5e 100644 --- a/src/topology/instances/add_circle.lean +++ b/src/topology/instances/add_circle.lean @@ -50,8 +50,58 @@ the rational circle `add_circle (1 : ℚ)`, and so we set things up more general noncomputable theory open set function add_subgroup topological_space +open_locale topological_space -variables {𝕜 : Type*} {B : Type*} +variables {𝕜 B : Type*} + +section continuity + +variables [linear_ordered_add_comm_group 𝕜] [archimedean 𝕜] + [topological_space 𝕜] [order_topology 𝕜] (a : 𝕜) {p : 𝕜} (hp : 0 < p) (x : 𝕜) + +lemma continuous_right_to_Ico_mod : continuous_within_at (to_Ico_mod a hp) (Ici x) x := +begin + intros s h, + rw [filter.mem_map, mem_nhds_within_iff_exists_mem_nhds_inter], + haveI : nontrivial 𝕜 := ⟨⟨0, p, hp.ne⟩⟩, + simp_rw mem_nhds_iff_exists_Ioo_subset at h ⊢, + obtain ⟨l, u, hxI, hIs⟩ := h, + let d := to_Ico_div a hp x • p, + have hd := to_Ico_mod_mem_Ico a hp x, + simp_rw [subset_def, mem_inter_iff], + refine ⟨_, ⟨l - d, min (a + p) u - d, _, λ x, id⟩, λ y, _⟩; + simp_rw [← add_mem_Ioo_iff_left, mem_Ioo, lt_min_iff], + { exact ⟨hxI.1, hd.2, hxI.2⟩ }, + { rintro ⟨h, h'⟩, apply hIs, + rw [← to_Ico_mod_add_zsmul, (to_Ico_mod_eq_self _).2], + exacts [⟨h.1, h.2.2⟩, ⟨hd.1.trans (add_le_add_right h' _), h.2.1⟩] }, +end + +lemma continuous_left_to_Ioc_mod : continuous_within_at (to_Ioc_mod a hp) (Iic x) x := +begin + rw (funext (λ y, eq.trans (by rw neg_neg) $ to_Ioc_mod_neg _ _ _) : + to_Ioc_mod a hp = (λ x, p - x) ∘ to_Ico_mod (-a) hp ∘ has_neg.neg), + exact ((continuous_sub_left _).continuous_at.comp_continuous_within_at $ + (continuous_right_to_Ico_mod _ _ _).comp continuous_neg.continuous_within_at $ λ y, neg_le_neg), +end + +variables {x} (hx : (x : 𝕜 ⧸ zmultiples p) ≠ a) + +lemma to_Ico_mod_eventually_eq_to_Ioc_mod : to_Ico_mod a hp =ᶠ[𝓝 x] to_Ioc_mod a hp := +is_open.mem_nhds (by {rw Ico_eq_locus_Ioc_eq_Union_Ioo, exact is_open_Union (λ i, is_open_Ioo)}) $ + (mem_Ioo_mod_iff_to_Ico_mod_eq_to_Ioc_mod hp).1 ((mem_Ioo_mod_iff_eq_mod_zmultiples hp).2 hx) + +lemma continuous_at_to_Ico_mod : continuous_at (to_Ico_mod a hp) x := +let h := to_Ico_mod_eventually_eq_to_Ioc_mod a hp hx in continuous_at_iff_continuous_left_right.2 $ + ⟨(continuous_left_to_Ioc_mod a hp x).congr_of_eventually_eq + (h.filter_mono nhds_within_le_nhds) h.eq_of_nhds, continuous_right_to_Ico_mod a hp x⟩ + +lemma continuous_at_to_Ioc_mod : continuous_at (to_Ioc_mod a hp) x := +let h := to_Ico_mod_eventually_eq_to_Ioc_mod a hp hx in continuous_at_iff_continuous_left_right.2 $ + ⟨continuous_left_to_Ioc_mod a hp x, (continuous_right_to_Ico_mod a hp x).congr_of_eventually_eq + (h.symm.filter_mono nhds_within_le_nhds) h.symm.eq_of_nhds⟩ + +end continuity /-- The "additive circle": `𝕜 ⧸ (ℤ ∙ p)`. See also `circle` and `real.angle`. -/ @[derive [add_comm_group, topological_space, topological_add_group, inhabited, has_coe_t 𝕜], @@ -93,7 +143,7 @@ end lemma coe_period : (p : add_circle p) = 0 := (quotient_add_group.eq_zero_iff p).2 $ mem_zmultiples p -@[simp] lemma coe_add_period (x : 𝕜) : (((x + p) : 𝕜) : add_circle p) = x := +@[simp] lemma coe_add_period (x : 𝕜) : ((x + p : 𝕜) : add_circle p) = x := by rw [coe_add, ←eq_sub_iff_add_eq', sub_self, coe_period] @[continuity, nolint unused_arguments] protected lemma continuous_mk' : @@ -105,12 +155,22 @@ include hp variables (a : 𝕜) [archimedean 𝕜] -/-- The natural equivalence between `add_circle p` and the half-open interval `[a, a + p)`. -/ +/-- The equivalence between `add_circle p` and the half-open interval `[a, a + p)`, whose inverse +is the natural quotient map. -/ def equiv_Ico : add_circle p ≃ Ico a (a + p) := quotient_add_group.equiv_Ico_mod a hp.out -/-- Given a function on `[a, a + p)`, lift it to `add_circle p`. -/ +/-- The equivalence between `add_circle p` and the half-open interval `(a, a + p]`, whose inverse +is the natural quotient map. -/ +def equiv_Ioc : add_circle p ≃ Ioc a (a + p) := quotient_add_group.equiv_Ioc_mod a hp.out + +/-- Given a function on `𝕜`, return the unique function on `add_circle p` agreeing with `f` on +`[a, a + p)`. -/ def lift_Ico (f : 𝕜 → B) : add_circle p → B := restrict _ f ∘ add_circle.equiv_Ico p a +/-- Given a function on `𝕜`, return the unique function on `add_circle p` agreeing with `f` on +`(a, a + p]`. -/ +def lift_Ioc (f : 𝕜 → B) : add_circle p → B := restrict _ f ∘ add_circle.equiv_Ioc p a + variables {p a} lemma coe_eq_coe_iff_of_mem_Ico {x y : 𝕜} @@ -132,16 +192,54 @@ begin refl, end +lemma lift_Ioc_coe_apply {f : 𝕜 → B} {x : 𝕜} (hx : x ∈ Ioc a (a + p)) : lift_Ioc p a f ↑x = f x := +begin + have : (equiv_Ioc p a) x = ⟨x, hx⟩, + { rw equiv.apply_eq_iff_eq_symm_apply, + refl, }, + rw [lift_Ioc, comp_apply, this], + refl, +end + variables (p a) +section continuity + @[continuity] lemma continuous_equiv_Ico_symm : continuous (equiv_Ico p a).symm := continuous_quotient_mk.comp continuous_subtype_coe +@[continuity] lemma continuous_equiv_Ioc_symm : continuous (equiv_Ioc p a).symm := +continuous_quotient_mk.comp continuous_subtype_coe + +variables {x : add_circle p} (hx : x ≠ a) +include hx + +lemma continuous_at_equiv_Ico : continuous_at (equiv_Ico p a) x := +begin + induction x using quotient_add_group.induction_on', + rw [continuous_at, filter.tendsto, quotient_add_group.nhds_eq, filter.map_map], + apply continuous_at.cod_restrict, exact continuous_at_to_Ico_mod a hp.out hx, +end + +lemma continuous_at_equiv_Ioc : continuous_at (equiv_Ioc p a) x := +begin + induction x using quotient_add_group.induction_on', + rw [continuous_at, filter.tendsto, quotient_add_group.nhds_eq, filter.map_map], + apply continuous_at.cod_restrict, exact continuous_at_to_Ioc_mod a hp.out hx, +end + +end continuity + /-- The image of the closed-open interval `[a, a + p)` under the quotient map `𝕜 → add_circle p` is the entire space. -/ @[simp] lemma coe_image_Ico_eq : (coe : 𝕜 → add_circle p) '' Ico a (a + p) = univ := by { rw image_eq_range, exact (equiv_Ico p a).symm.range_eq_univ } +/-- The image of the closed-open interval `[a, a + p)` under the quotient map `𝕜 → add_circle p` is +the entire space. -/ +@[simp] lemma coe_image_Ioc_eq : (coe : 𝕜 → add_circle p) '' Ioc a (a + p) = univ := +by { rw image_eq_range, exact (equiv_Ioc p a).symm.range_eq_univ } + /-- The image of the closed interval `[0, p]` under the quotient map `𝕜 → add_circle p` is the entire space. -/ @[simp] lemma coe_image_Icc_eq : (coe : 𝕜 → add_circle p) '' Icc a (a + p) = univ := @@ -357,14 +455,12 @@ local attribute [instance] fact_zero_lt_one abbreviation unit_add_circle := add_circle (1 : ℝ) section identify_Icc_ends -/-! This section proves that for any `a`, the natural map from `[a, a + p] ⊂ ℝ` to `add_circle p` +/-! This section proves that for any `a`, the natural map from `[a, a + p] ⊂ 𝕜` to `add_circle p` gives an identification of `add_circle p`, as a topological space, with the quotient of `[a, a + p]` by the equivalence relation identifying the endpoints. -/ namespace add_circle -section linear_ordered_add_comm_group - variables [linear_ordered_add_comm_group 𝕜] [topological_space 𝕜] [order_topology 𝕜] (p a : 𝕜) [hp : fact (0 < p)] @@ -383,7 +479,7 @@ variables [archimedean 𝕜] /-- The equivalence between `add_circle p` and the quotient of `[a, a + p]` by the relation identifying the endpoints. -/ def equiv_Icc_quot : 𝕋 ≃ quot (endpoint_ident p a) := -{ to_fun := λ x, quot.mk _ $ subtype.map id Ico_subset_Icc_self (equiv_Ico _ _ x), +{ to_fun := λ x, quot.mk _ $ inclusion Ico_subset_Icc_self (equiv_Ico _ _ x), inv_fun := λ x, quot.lift_on x coe $ by { rintro _ _ ⟨_⟩, exact (coe_add_period p a).symm }, left_inv := (equiv_Ico p a).symm_apply_apply, right_inv := quot.ind $ by @@ -396,54 +492,61 @@ def equiv_Icc_quot : 𝕋 ≃ quot (endpoint_ident p a) := congr, ext1, apply congr_arg subtype.val ((equiv_Ico p a).right_inv ⟨x, hx.1, hx.2.lt_of_ne h⟩) } } -end linear_ordered_add_comm_group - -section real +lemma equiv_Icc_quot_comp_mk_eq_to_Ico_mod : equiv_Icc_quot p a ∘ quotient.mk' = + λ x, quot.mk _ ⟨to_Ico_mod a hp.out x, Ico_subset_Icc_self $ to_Ico_mod_mem_Ico a _ x⟩ := rfl -variables {p a : ℝ} [hp : fact (0 < p)] -include hp - -local notation `𝕋` := add_circle p - -/- doesn't work if inlined in `homeo_of_equiv_compact_to_t2` -- why? -/ -private lemma continuous_equiv_Icc_quot_symm : continuous (equiv_Icc_quot p a).symm := -continuous_quot_lift _ $ (add_circle.continuous_mk' p).comp continuous_subtype_coe +lemma equiv_Icc_quot_comp_mk_eq_to_Ioc_mod : equiv_Icc_quot p a ∘ quotient.mk' = + λ x, quot.mk _ ⟨to_Ioc_mod a hp.out x, Ioc_subset_Icc_self $ to_Ioc_mod_mem_Ioc a _ x⟩ := +begin + rw equiv_Icc_quot_comp_mk_eq_to_Ico_mod, funext, + by_cases mem_Ioo_mod a p x, + { simp_rw (mem_Ioo_mod_iff_to_Ico_mod_eq_to_Ioc_mod hp.out).1 h }, + { simp_rw [not_imp_comm.1 (mem_Ioo_mod_iff_to_Ico_mod_ne_left hp.out).2 h, + not_imp_comm.1 (mem_Ioo_mod_iff_to_Ioc_mod_ne_right hp.out).2 h], + exact quot.sound endpoint_ident.mk }, +end -/-- The natural map from `[a, a + p] ⊂ ℝ` with endpoints identified to `ℝ / ℤ • p`, as a +/-- The natural map from `[a, a + p] ⊂ 𝕜` with endpoints identified to `𝕜 / ℤ • p`, as a homeomorphism of topological spaces. -/ -def homeo_Icc_quot : 𝕋 ≃ₜ quot (endpoint_ident p a) := -(continuous.homeo_of_equiv_compact_to_t2 continuous_equiv_Icc_quot_symm).symm +def homeo_Icc_quot : 𝕋 ≃ₜ quot (endpoint_ident p a) := +{ to_equiv := equiv_Icc_quot p a, + continuous_to_fun := begin + simp_rw [quotient_map_quotient_mk.continuous_iff, + continuous_iff_continuous_at, continuous_at_iff_continuous_left_right], + intro x, split, + work_on_goal 1 { erw equiv_Icc_quot_comp_mk_eq_to_Ioc_mod }, + work_on_goal 2 { erw equiv_Icc_quot_comp_mk_eq_to_Ico_mod }, + all_goals { apply continuous_quot_mk.continuous_at.comp_continuous_within_at, + rw inducing_coe.continuous_within_at_iff }, + { apply continuous_left_to_Ioc_mod }, + { apply continuous_right_to_Ico_mod }, + end, + continuous_inv_fun := continuous_quot_lift _ + ((add_circle.continuous_mk' p).comp continuous_subtype_coe) } + +/-! We now show that a continuous function on `[a, a + p]` satisfying `f a = f (a + p)` is the +pullback of a continuous function on `add_circle p`. -/ -/-! We now show that a continuous function on `[a, a + p]` satisfying `f a = f (a + p)` is -the pullback of a continuous function on `unit_add_circle`. -/ - -lemma eq_of_end_ident {f : ℝ → B} (hf : f a = f (a + p)) (x y : Icc a (a + p)) : - endpoint_ident p a x y → f x = f y := by { rintro ⟨_⟩, exact hf } +variables {p a} -lemma lift_Ico_eq_lift_Icc {f : ℝ → B} (h : f a = f (a + p)) : - lift_Ico p a f = (quot.lift (restrict (Icc a $ a + p) f) $ eq_of_end_ident h) - ∘ equiv_Icc_quot p a := -funext (λ x, by refl) +lemma lift_Ico_eq_lift_Icc {f : 𝕜 → B} (h : f a = f (a + p)) : lift_Ico p a f = + quot.lift (restrict (Icc a $ a + p) f) (by { rintro _ _ ⟨_⟩, exact h }) ∘ equiv_Icc_quot p a := +rfl -lemma lift_Ico_continuous [topological_space B] {f : ℝ → B} (hf : f a = f (a + p)) +lemma lift_Ico_continuous [topological_space B] {f : 𝕜 → B} (hf : f a = f (a + p)) (hc : continuous_on f $ Icc a (a + p)) : continuous (lift_Ico p a f) := begin rw lift_Ico_eq_lift_Icc hf, - refine continuous.comp _ homeo_Icc_quot.continuous_to_fun, + refine continuous.comp _ (homeo_Icc_quot p a).continuous_to_fun, exact continuous_coinduced_dom.mpr (continuous_on_iff_continuous_restrict.mp hc), end -end real - section zero_based -variables {p : ℝ} [hp : fact (0 < p)] -include hp - -lemma lift_Ico_zero_coe_apply {f : ℝ → B} {x : ℝ} (hx : x ∈ Ico 0 p) : +lemma lift_Ico_zero_coe_apply {f : 𝕜 → B} {x : 𝕜} (hx : x ∈ Ico 0 p) : lift_Ico p 0 f ↑x = f x := lift_Ico_coe_apply (by rwa zero_add) -lemma lift_Ico_zero_continuous [topological_space B] {f : ℝ → B} +lemma lift_Ico_zero_continuous [topological_space B] {f : 𝕜 → B} (hf : f 0 = f p) (hc : continuous_on f $ Icc 0 p) : continuous (lift_Ico p 0 f) := lift_Ico_continuous (by rwa zero_add : f 0 = f (0 + p)) (by rwa zero_add) From 5a1fa97e87785d3d0b268769f71076ccec900fa2 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sun, 8 Jan 2023 00:31:05 +0000 Subject: [PATCH 0175/1291] chore(number_theory/pell): golf, use tactic mode (#18091) Also drop an unneeded assumption in `pell.eq_pow_of_pell_lem`. --- src/number_theory/pell.lean | 196 +++++++++++++++--------------------- 1 file changed, 83 insertions(+), 113 deletions(-) diff --git a/src/number_theory/pell.lean b/src/number_theory/pell.lean index 2ce2c8408b372..5fed9482d43e1 100644 --- a/src/number_theory/pell.lean +++ b/src/number_theory/pell.lean @@ -683,119 +683,89 @@ theorem matiyasevic {a k x y} : (∃ a1 : 1 < a, xn a1 k = x ∧ yn a1 k = y) end end⟩⟩ -lemma eq_pow_of_pell_lem {a y k} (a1 : 1 < a) (ypos : 0 < y) : 0 < k → y^k < a → +lemma eq_pow_of_pell_lem {a y k} (hy0 : y ≠ 0) (hk0 : k ≠ 0) (hyk : y^k < a) : (↑(y^k) : ℤ) < 2*a*y - y*y - 1 := -have y < a → a + (y*y + 1) ≤ 2*a*y, begin - intro ya, induction y with y IH, exact absurd ypos (lt_irrefl _), - cases nat.eq_zero_or_pos y with y0 ypos, - { rw y0, simpa [two_mul], }, - { rw [nat.mul_succ, nat.mul_succ, nat.succ_mul y], - have : y + nat.succ y ≤ 2 * a, - { change y + y < 2 * a, rw ← two_mul, - exact mul_lt_mul_of_pos_left (nat.lt_of_succ_lt ya) dec_trivial }, - have := add_le_add (IH ypos (nat.lt_of_succ_lt ya)) this, - convert this using 1, - ring } -end, λk0 yak, -lt_of_lt_of_le (int.coe_nat_lt_coe_nat_of_lt yak) $ -by rw sub_sub; apply le_sub_right_of_add_le; - apply int.coe_nat_le_coe_nat_of_le; - have y1 := nat.pow_le_pow_of_le_right ypos k0; simp at y1; - exact this (lt_of_le_of_lt y1 yak) - -theorem eq_pow_of_pell {m n k} : (n^k = m ↔ -k = 0 ∧ m = 1 ∨ 0 < k ∧ -(n = 0 ∧ m = 0 ∨ 0 < n ∧ -∃ (w a t z : ℕ) (a1 : 1 < a), - xn a1 k ≡ yn a1 k * (a - n) + m [MOD t] ∧ - 2 * a * n = t + (n * n + 1) ∧ - m < t ∧ n ≤ w ∧ k ≤ w ∧ - a * a - ((w + 1) * (w + 1) - 1) * (w * z) * (w * z) = 1)) := -⟨λe, by rw ← e; - refine (nat.eq_zero_or_pos k).elim - (λk0, by rw k0; exact or.inl ⟨rfl, rfl⟩) - (λkpos, or.inr ⟨kpos, _⟩); - refine (nat.eq_zero_or_pos n).elim - (λn0, by rw [n0, zero_pow kpos]; exact or.inl ⟨rfl, rfl⟩) - (λnpos, or.inr ⟨npos, _⟩); exact - let w := max n k in - have nw : n ≤ w, from le_max_left _ _, - have kw : k ≤ w, from le_max_right _ _, - have wpos : 0 < w, from lt_of_lt_of_le npos nw, - have w1 : 1 < w + 1, from nat.succ_lt_succ wpos, - let a := xn w1 w in - have a1 : 1 < a, from strict_mono_x w1 wpos, - let x := xn a1 k, y := yn a1 k in - let ⟨z, ze⟩ := show w ∣ yn w1 w, from modeq_zero_iff_dvd.1 $ - (yn_modeq_a_sub_one w1 w).trans dvd_rfl.modeq_zero_nat in - have nt : (↑(n^k) : ℤ) < 2 * a * n - n * n - 1, from - eq_pow_of_pell_lem a1 npos kpos $ calc - n^k ≤ n^w : nat.pow_le_pow_of_le_right npos kw - ... < (w + 1)^w : nat.pow_lt_pow_of_lt_left (nat.lt_succ_of_le nw) wpos - ... ≤ a : xn_ge_a_pow w1 w, - let ⟨t, te⟩ := int.eq_coe_of_zero_le $ - le_trans (int.coe_zero_le _) nt.le in - have na : n ≤ a, from nw.trans $ le_of_lt $ n_lt_xn w1 w, - have tm : x ≡ y * (a - n) + n^k [MOD t], begin - apply modeq_of_dvd, - rw [int.coe_nat_add, int.coe_nat_mul, int.coe_nat_sub na, ← te], - exact x_sub_y_dvd_pow a1 n k - end, - have ta : 2 * a * n = t + (n * n + 1), from int.coe_nat_inj $ - by rw [int.coe_nat_add, ← te, sub_sub]; - repeat {rw int.coe_nat_add <|> rw int.coe_nat_mul}; - rw [int.coe_nat_one, sub_add_cancel]; refl, - have mt : n^k < t, from int.lt_of_coe_nat_lt_coe_nat $ - by rw ← te; exact nt, - have zp : a * a - ((w + 1) * (w + 1) - 1) * (w * z) * (w * z) = 1, - by rw ← ze; exact pell_eq w1 w, - ⟨w, a, t, z, a1, tm, ta, mt, nw, kw, zp⟩, -λo, match o with -| or.inl ⟨k0, m1⟩ := by rw [k0, m1]; refl -| or.inr ⟨kpos, or.inl ⟨n0, m0⟩⟩ := by rw [n0, m0, zero_pow kpos] -| or.inr ⟨kpos, or.inr ⟨npos, w, a, t, z, - (a1 : 1 < a), - (tm : xn a1 k ≡ yn a1 k * (a - n) + m [MOD t]), - (ta : 2 * a * n = t + (n * n + 1)), - (mt : m < t), - (nw : n ≤ w), - (kw : k ≤ w), - (zp : a * a - ((w + 1) * (w + 1) - 1) * (w * z) * (w * z) = 1)⟩⟩ := - have wpos : 0 < w, from lt_of_lt_of_le npos nw, - have w1 : 1 < w + 1, from nat.succ_lt_succ wpos, - let ⟨j, xj, yj⟩ := eq_pell w1 zp in - by clear _match o _let_match; exact - have jpos : 0 < j, from (nat.eq_zero_or_pos j).resolve_left $ λj0, - have a1 : a = 1, by rw j0 at xj; exact xj, - have 2 * n = t + (n * n + 1), by rw a1 at ta; exact ta, - have n1 : n = 1, from - have n * n < n * 2, by rw [mul_comm n 2, this]; apply nat.le_add_left, - have n ≤ 1, from nat.le_of_lt_succ $ lt_of_mul_lt_mul_left this (nat.zero_le _), - le_antisymm this npos, - by rw n1 at this; - rw ← @nat.add_right_cancel 0 2 t this at mt; - exact nat.not_lt_zero _ mt, - have wj : w ≤ j, from nat.le_of_dvd jpos $ modeq_zero_iff_dvd.1 $ - (yn_modeq_a_sub_one w1 j).symm.trans $ - modeq_zero_iff_dvd.2 ⟨z, yj.symm⟩, - have nt : (↑(n^k) : ℤ) < 2 * a * n - n * n - 1, from - eq_pow_of_pell_lem a1 npos kpos $ calc - n^k ≤ n^j : nat.pow_le_pow_of_le_right npos (le_trans kw wj) - ... < (w + 1)^j : nat.pow_lt_pow_of_lt_left (nat.lt_succ_of_le nw) jpos - ... ≤ xn w1 j : xn_ge_a_pow w1 j - ... = a : xj.symm, - have na : n ≤ a, by rw xj; exact - le_trans (le_trans nw wj) (le_of_lt $ n_lt_xn _ _), - have te : (t : ℤ) = 2 * ↑a * ↑n - ↑n * ↑n - 1, by - rw sub_sub; apply eq_sub_of_add_eq; apply (int.coe_nat_eq_coe_nat_iff _ _).2; - exact ta.symm, - have xn a1 k ≡ yn a1 k * (a - n) + n^k [MOD t], - by have := x_sub_y_dvd_pow a1 n k; - rw [← te, ← int.coe_nat_sub na] at this; exact modeq_of_dvd this, - have n^k % t = m % t, from - (this.symm.trans tm).add_left_cancel' _, - by rw ← te at nt; - rwa [nat.mod_eq_of_lt (int.lt_of_coe_nat_lt_coe_nat nt), nat.mod_eq_of_lt mt] at this -end⟩ +have hya : y < a, from (nat.le_self_pow hk0 _).trans_lt hyk, +calc (↑(y ^ k) : ℤ) < a : nat.cast_lt.2 hyk +... ≤ a ^ 2 - (a - 1) ^ 2 - 1 : + begin + rw [sub_sq, mul_one, one_pow, sub_add, sub_sub_cancel, two_mul, sub_sub, ← add_sub, + le_add_iff_nonneg_right, ← bit0, sub_nonneg, ← nat.cast_two, nat.cast_le, nat.succ_le_iff], + exact (one_le_iff_ne_zero.2 hy0).trans_lt hya + end +... ≤ a ^ 2 - (a - y) ^ 2 - 1 : have _ := hya.le, + by { mono*; simpa only [sub_nonneg, nat.cast_le, nat.one_le_cast, nat.one_le_iff_ne_zero] } +... = 2*a*y - y*y - 1 : by ring + +theorem eq_pow_of_pell {m n k} : n^k = m ↔ + k = 0 ∧ m = 1 ∨ + 0 < k ∧ (n = 0 ∧ m = 0 ∨ + 0 < n ∧ ∃ (w a t z : ℕ) (a1 : 1 < a), + xn a1 k ≡ yn a1 k * (a - n) + m [MOD t] ∧ + 2 * a * n = t + (n * n + 1) ∧ + m < t ∧ n ≤ w ∧ k ≤ w ∧ + a * a - ((w + 1) * (w + 1) - 1) * (w * z) * (w * z) = 1) := +begin + split, + { rintro rfl, + refine k.eq_zero_or_pos.imp (λ k0, k0.symm ▸ ⟨rfl, rfl⟩) (λ hk, ⟨hk, _⟩), + refine n.eq_zero_or_pos.imp (λ n0, n0.symm ▸ ⟨rfl, zero_pow hk⟩) (λ hn, ⟨hn, _⟩), + set w := max n k, + have nw : n ≤ w, from le_max_left _ _, + have kw : k ≤ w, from le_max_right _ _, + have wpos : 0 < w, from hn.trans_le nw, + have w1 : 1 < w + 1, from nat.succ_lt_succ wpos, + set a := xn w1 w, + have a1 : 1 < a, from strict_mono_x w1 wpos, + have na : n ≤ a, from nw.trans (n_lt_xn w1 w).le, + set x := xn a1 k, set y := yn a1 k, + obtain ⟨z, ze⟩ : w ∣ yn w1 w, + from modeq_zero_iff_dvd.1 ((yn_modeq_a_sub_one w1 w).trans dvd_rfl.modeq_zero_nat), + have nt : (↑(n^k) : ℤ) < 2 * a * n - n * n - 1, + { refine eq_pow_of_pell_lem hn.ne' hk.ne' _, + calc n^k ≤ n^w : nat.pow_le_pow_of_le_right hn kw + ... < (w + 1)^w : nat.pow_lt_pow_of_lt_left (nat.lt_succ_of_le nw) wpos + ... ≤ a : xn_ge_a_pow w1 w }, + lift (2 * a * n - n * n - 1 : ℤ) to ℕ using ((nat.cast_nonneg _).trans nt.le) with t te, + have tm : x ≡ y * (a - n) + n^k [MOD t], + { apply modeq_of_dvd, + rw [int.coe_nat_add, int.coe_nat_mul, int.coe_nat_sub na, te], + exact x_sub_y_dvd_pow a1 n k }, + have ta : 2 * a * n = t + (n * n + 1), + { rw [← @nat.cast_inj ℤ, int.coe_nat_add, te, sub_sub], + repeat { rw nat.cast_add <|> rw nat.cast_mul }, + rw [nat.cast_one, sub_add_cancel, nat.cast_two] }, + have zp : a * a - ((w + 1) * (w + 1) - 1) * (w * z) * (w * z) = 1, + from ze ▸ pell_eq w1 w, + exact ⟨w, a, t, z, a1, tm, ta, nat.cast_lt.1 nt, nw, kw, zp⟩ }, + { rintro (⟨rfl, rfl⟩ | ⟨hk0, ⟨rfl, rfl⟩ | ⟨hn0, w, a, t, z, a1, tm, ta, mt, nw, kw, zp⟩⟩), + { exact pow_zero n }, { exact zero_pow hk0 }, + have hw0 : 0 < w, from hn0.trans_le nw, + have hw1 : 1 < w + 1, from nat.succ_lt_succ hw0, + rcases eq_pell hw1 zp with ⟨j, rfl, yj⟩, + have hj0 : 0 < j, + { apply nat.pos_of_ne_zero, + rintro rfl, + exact lt_irrefl 1 a1 }, + have wj : w ≤ j := nat.le_of_dvd hj0 (modeq_zero_iff_dvd.1 $ + (yn_modeq_a_sub_one hw1 j).symm.trans $ modeq_zero_iff_dvd.2 ⟨z, yj.symm⟩), + have hnka : n ^ k < xn hw1 j, + calc n^k ≤ n^j : nat.pow_le_pow_of_le_right hn0 (le_trans kw wj) + ... < (w + 1)^j : nat.pow_lt_pow_of_lt_left (nat.lt_succ_of_le nw) hj0 + ... ≤ xn hw1 j : xn_ge_a_pow hw1 j, + have nt : (↑(n^k) : ℤ) < 2 * xn hw1 j * n - n * n - 1, + from eq_pow_of_pell_lem hn0.ne' hk0.ne' hnka, + have na : n ≤ xn hw1 j, from (nat.le_self_pow hk0.ne' _).trans hnka.le, + have te : (t : ℤ) = 2 * xn hw1 j * n - n * n - 1, + { rw [sub_sub, eq_sub_iff_add_eq], + exact_mod_cast ta.symm }, + have : xn a1 k ≡ yn a1 k * (xn hw1 j - n) + n^k [MOD t], + { apply modeq_of_dvd, + rw [te, nat.cast_add, nat.cast_mul, int.coe_nat_sub na], + exact x_sub_y_dvd_pow a1 n k }, + have : n^k % t = m % t, from (this.symm.trans tm).add_left_cancel' _, + rw [← te] at nt, + rwa [nat.mod_eq_of_lt (nat.cast_lt.1 nt), nat.mod_eq_of_lt mt] at this } +end end pell From e7bab9a85e92cf46c02cb4725a7be2f04691e3a7 Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Sun, 8 Jan 2023 05:18:21 +0000 Subject: [PATCH 0176/1291] chore(algebra/group_ring_action/invariant): streamline imports (#18092) The only file importing `algebra/group_ring_action/invariant` was `algebra/hom/group_action`. This means that the material needing both files can safely be moved from `hom/group_action` to `group_ring_action/invariant`, while only decreasing the imports of anything else in the hierarchy. This is worth doing since `group_ring_action/invariant` has another relatively heavy import (`ring_theory/subring/pointwise`). After this rearrangement, `hom/group_action` should be ready to port. --- src/algebra/algebra/basic.lean | 1 + src/algebra/group_ring_action/invariant.lean | 18 +++++++++++++++ src/algebra/hom/group_action.lean | 23 ++++---------------- src/algebra/module/linear_map.lean | 1 + src/field_theory/fixed.lean | 1 + src/ring_theory/ideal/quotient.lean | 1 + 6 files changed, 26 insertions(+), 19 deletions(-) diff --git a/src/algebra/algebra/basic.lean b/src/algebra/algebra/basic.lean index a20c4ffa7eb78..b478de90292dc 100644 --- a/src/algebra/algebra/basic.lean +++ b/src/algebra/algebra/basic.lean @@ -10,6 +10,7 @@ import algebra.ring.aut import algebra.ring.ulift import algebra.char_zero.lemmas import linear_algebra.span +import ring_theory.subring.basic import tactic.abel /-! diff --git a/src/algebra/group_ring_action/invariant.lean b/src/algebra/group_ring_action/invariant.lean index 3e10bde74ee2b..f0f9a7d16a4a6 100644 --- a/src/algebra/group_ring_action/invariant.lean +++ b/src/algebra/group_ring_action/invariant.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Eric Wieser. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Eric Wieser -/ +import algebra.hom.group_action import ring_theory.subring.pointwise /-! # Subrings invariant under an action -/ @@ -28,3 +29,20 @@ instance is_invariant_subring.to_mul_semiring_action [is_invariant_subring M S] smul_mul := λ m s₁ s₂, subtype.eq $ smul_mul' m s₁ s₂ } end ring + +section +variables (M : Type*) [monoid M] +variables {R' : Type*} [ring R'] [mul_semiring_action M R'] +variables (U : subring R') [is_invariant_subring M U] + +/-- The canonical inclusion from an invariant subring. -/ +def is_invariant_subring.subtype_hom : U →+*[M] R' := +{ map_smul' := λ m s, rfl, ..U.subtype } + +@[simp] theorem is_invariant_subring.coe_subtype_hom : + (is_invariant_subring.subtype_hom M U : U → R') = coe := rfl + +@[simp] theorem is_invariant_subring.coe_subtype_hom' : + (is_invariant_subring.subtype_hom M U : U →+* R') = U.subtype := rfl + +end diff --git a/src/algebra/hom/group_action.lean b/src/algebra/hom/group_action.lean index 4bc204da1c7d4..a2e4343592e30 100644 --- a/src/algebra/hom/group_action.lean +++ b/src/algebra/hom/group_action.lean @@ -3,9 +3,8 @@ Copyright (c) 2020 Kenny Lau. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau -/ -import algebra.group_ring_action.invariant -import group_theory.group_action.defs -import group_theory.subgroup.basic +import algebra.group_ring_action.basic +import algebra.module.basic /-! # Equivariant homomorphisms @@ -35,6 +34,8 @@ The above types have corresponding classes: -/ +assert_not_exists submonoid + variables (M' : Type*) variables (X : Type*) [has_smul M' X] variables (Y : Type*) [has_smul M' Y] @@ -50,7 +51,6 @@ variables (R' : Type*) [ring R'] [mul_semiring_action M R'] variables (S : Type*) [semiring S] [mul_semiring_action M S] variables (S' : Type*) [ring S'] [mul_semiring_action M S'] variables (T : Type*) [semiring T] [mul_semiring_action M T] -variables (G : Type*) [group G] (H : subgroup G) set_option old_structure_cmd true @@ -340,18 +340,3 @@ ext $ λ x, by rw [comp_apply, id_apply] ext $ λ x, by rw [comp_apply, id_apply] end mul_semiring_action_hom - -section -variables (M) {R'} (U : subring R') [is_invariant_subring M U] - -/-- The canonical inclusion from an invariant subring. -/ -def is_invariant_subring.subtype_hom : U →+*[M] R' := -{ map_smul' := λ m s, rfl, ..U.subtype } - -@[simp] theorem is_invariant_subring.coe_subtype_hom : - (is_invariant_subring.subtype_hom M U : U → R') = coe := rfl - -@[simp] theorem is_invariant_subring.coe_subtype_hom' : - (is_invariant_subring.subtype_hom M U : U →+* R') = U.subtype := rfl - -end diff --git a/src/algebra/module/linear_map.lean b/src/algebra/module/linear_map.lean index 96457fd912eb8..affad643cf562 100644 --- a/src/algebra/module/linear_map.lean +++ b/src/algebra/module/linear_map.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Nathaniel Thomas, Jeremy Avigad, Johannes Hölzl, Mario Carneiro, Anne Baanen, Frédéric Dupuis, Heather Macbeth -/ +import algebra.big_operators.basic import algebra.hom.group_action import algebra.module.pi import algebra.star.basic diff --git a/src/field_theory/fixed.lean b/src/field_theory/fixed.lean index 599e38b164127..f8192a0368438 100644 --- a/src/field_theory/fixed.lean +++ b/src/field_theory/fixed.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau -/ +import algebra.group_ring_action.invariant import algebra.polynomial.group_ring_action import field_theory.normal import field_theory.separable diff --git a/src/ring_theory/ideal/quotient.lean b/src/ring_theory/ideal/quotient.lean index a1b0740ec5350..ce7a83a3f0eae 100644 --- a/src/ring_theory/ideal/quotient.lean +++ b/src/ring_theory/ideal/quotient.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau, Chris Hughes, Mario Carneiro, Anne Baanen -/ import algebra.ring.fin +import algebra.ring.prod import linear_algebra.quotient import ring_theory.congruence import ring_theory.ideal.basic From e001509c11c4d0f549d91d89da95b4a0b43c714f Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sun, 8 Jan 2023 08:40:59 +0000 Subject: [PATCH 0177/1291] feat(field_theory/chevalley_warning): Binary version (#18083) Derive the two multivariate polynomials version of Chevalley-Warning from the general one. --- src/field_theory/chevalley_warning.lean | 58 +++++++++++++++++-------- 1 file changed, 39 insertions(+), 19 deletions(-) diff --git a/src/field_theory/chevalley_warning.lean b/src/field_theory/chevalley_warning.lean index cad0f2899a95b..3e0de4a8b463b 100644 --- a/src/field_theory/chevalley_warning.lean +++ b/src/field_theory/chevalley_warning.lean @@ -18,8 +18,8 @@ and `q` is notation for the cardinality of `K`. 1. Let `f` be a multivariate polynomial in finitely many variables (`X s`, `s : σ`) such that the total degree of `f` is less than `(q-1)` times the cardinality of `σ`. Then the evaluation of `f` on all points of `σ → K` (aka `K^σ`) sums to `0`. - (`sum_mv_polynomial_eq_zero`) -2. The Chevalley–Warning theorem (`char_dvd_card_solutions`). + (`sum_eval_eq_zero`) +2. The Chevalley–Warning theorem (`char_dvd_card_solutions_of_sum_lt`). Let `f i` be a finite family of multivariate polynomials in finitely many variables (`X s`, `s : σ`) such that the sum of the total degrees of the `f i` is less than the cardinality of `σ`. @@ -41,12 +41,12 @@ open_locale big_operators section finite_field open mv_polynomial function (hiding eval) finset finite_field -variables {K : Type*} {σ : Type*} [fintype K] [field K] [fintype σ] +variables {K σ ι : Type*} [fintype K] [field K] [fintype σ] [decidable_eq σ] local notation `q` := fintype.card K -lemma mv_polynomial.sum_mv_polynomial_eq_zero [decidable_eq σ] (f : mv_polynomial σ K) +lemma mv_polynomial.sum_eval_eq_zero (f : mv_polynomial σ K) (h : f.total_degree < (q - 1) * fintype.card σ) : - (∑ x, eval x f) = 0 := + ∑ x, eval x f = 0 := begin haveI : decidable_eq K := classical.dec_eq K, calc (∑ x, eval x f) @@ -86,15 +86,14 @@ begin rw equiv.subtype_equiv_codomain_symm_apply_ne, } end -variables [decidable_eq K] [decidable_eq σ] +variables [decidable_eq K] (p : ℕ) [char_p K p] -/-- The Chevalley–Warning theorem. +/-- The **Chevalley–Warning theorem**, finitary version. Let `(f i)` be a finite family of multivariate polynomials in finitely many variables (`X s`, `s : σ`) over a finite field of characteristic `p`. Assume that the sum of the total degrees of the `f i` is less than the cardinality of `σ`. Then the number of common solutions of the `f i` is divisible by `p`. -/ -theorem char_dvd_card_solutions_family (p : ℕ) [char_p K p] - {ι : Type*} {s : finset ι} {f : ι → mv_polynomial σ K} +theorem char_dvd_card_solutions_of_sum_lt {s : finset ι} {f : ι → mv_polynomial σ K} (h : (∑ i in s, (f i).total_degree) < fintype.card σ) : p ∣ fintype.card {x : σ → K // ∀ i ∈ s, eval x (f i) = 0} := begin @@ -131,7 +130,7 @@ begin rw [← char_p.cast_eq_zero_iff K, ← key], show ∑ x, eval x F = 0, -- We are now ready to apply the main machine, proven before. - apply F.sum_mv_polynomial_eq_zero, + apply F.sum_eval_eq_zero, -- It remains to verify the crucial assumption of this machine show F.total_degree < (q - 1) * fintype.card σ, calc F.total_degree ≤ ∑ i in s, (1 - (f i)^(q - 1)).total_degree : total_degree_finset_prod s _ @@ -147,22 +146,43 @@ begin ... ≤ (q - 1) * (f i).total_degree : total_degree_pow _ _ end -/-- The Chevalley–Warning theorem. +/-- The **Chevalley–Warning theorem**, fintype version. +Let `(f i)` be a finite family of multivariate polynomials +in finitely many variables (`X s`, `s : σ`) over a finite field of characteristic `p`. +Assume that the sum of the total degrees of the `f i` is less than the cardinality of `σ`. +Then the number of common solutions of the `f i` is divisible by `p`. -/ +theorem char_dvd_card_solutions_of_fintype_sum_lt [fintype ι] {f : ι → mv_polynomial σ K} + (h : (∑ i, (f i).total_degree) < fintype.card σ) : + p ∣ fintype.card {x : σ → K // ∀ i, eval x (f i) = 0} := +by simpa using char_dvd_card_solutions_of_sum_lt p h + +/-- The **Chevalley–Warning theorem**, unary version. Let `f` be a multivariate polynomial in finitely many variables (`X s`, `s : σ`) over a finite field of characteristic `p`. Assume that the total degree of `f` is less than the cardinality of `σ`. Then the number of solutions of `f` is divisible by `p`. -See `char_dvd_card_solutions_family` for a version that takes a family of polynomials `f i`. -/ -theorem char_dvd_card_solutions (p : ℕ) [char_p K p] - {f : mv_polynomial σ K} (h : f.total_degree < fintype.card σ) : +See `char_dvd_card_solutions_of_sum_lt` for a version that takes a family of polynomials `f i`. -/ +theorem char_dvd_card_solutions {f : mv_polynomial σ K} (h : f.total_degree < fintype.card σ) : p ∣ fintype.card {x : σ → K // eval x f = 0} := begin let F : unit → mv_polynomial σ K := λ _, f, - have : ∑ i : unit, (F i).total_degree < fintype.card σ, - { simpa only [fintype.univ_punit, sum_singleton] using h, }, - have key := char_dvd_card_solutions_family p this, - simp only [F, fintype.univ_punit, forall_eq, mem_singleton] at key, - convert key, + have : ∑ i : unit, (F i).total_degree < fintype.card σ := h, + simpa only [F, fintype.univ_punit, forall_eq, mem_singleton] using + char_dvd_card_solutions_of_sum_lt p this, +end + +/-- The **Chevalley–Warning theorem**, binary version. +Let `f₁`, `f₂` be two multivariate polynomials in finitely many variables (`X s`, `s : σ`) over a +finite field of characteristic `p`. +Assume that the sum of the total degrees of `f₁` and `f₂` is less than the cardinality of `σ`. +Then the number of common solutions of the `f₁` and `f₂` is divisible by `p`. -/ +theorem char_dvd_card_solutions_of_add_lt {f₁ f₂ : mv_polynomial σ K} + (h : f₁.total_degree + f₂.total_degree < fintype.card σ) : + p ∣ fintype.card {x : σ → K // eval x f₁ = 0 ∧ eval x f₂ = 0} := +begin + let F : bool → mv_polynomial σ K := λ b, cond b f₂ f₁, + have : ∑ b : bool, (F b).total_degree < fintype.card σ := (add_comm _ _).trans_lt h, + simpa only [F, bool.forall_bool] using char_dvd_card_solutions_of_fintype_sum_lt p this, end end finite_field From bc7d81beddb3d6c66f71449c5bc76c38cb77cf9e Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sun, 8 Jan 2023 09:48:24 +0000 Subject: [PATCH 0178/1291] feat(data/{set,finset}/basic): Convenience lemmas (#17957) Add a few convenient corollaries to existing lemmas. --- counterexamples/phillips.lean | 8 +- src/algebra/indicator_function.lean | 4 +- .../box_integral/partition/basic.lean | 2 +- .../partition/subbox_induction.lean | 2 +- src/data/finset/basic.lean | 44 ++++-- src/data/finset/card.lean | 12 ++ src/data/finset/image.lean | 3 + src/data/finset/lattice.lean | 8 +- src/data/fintype/basic.lean | 21 ++- src/data/set/basic.lean | 130 ++++++++---------- src/data/set/lattice.lean | 33 +++-- src/group_theory/perm/cycle/type.lean | 3 +- src/measure_theory/decomposition/jordan.lean | 4 +- .../decomposition/signed_hahn.lean | 4 +- .../decomposition/unsigned_hahn.lean | 7 +- src/measure_theory/integral/set_integral.lean | 4 +- src/measure_theory/measure/ae_disjoint.lean | 3 +- src/measure_theory/measure/measure_space.lean | 4 +- src/measure_theory/measure/outer_measure.lean | 2 +- .../measure/vector_measure.lean | 8 +- src/order/boolean_algebra.lean | 7 + src/order/rel_classes.lean | 10 +- src/topology/instances/ennreal.lean | 2 +- 23 files changed, 190 insertions(+), 135 deletions(-) diff --git a/counterexamples/phillips.lean b/counterexamples/phillips.lean index 1be7228367598..d5221ec1e4383 100644 --- a/counterexamples/phillips.lean +++ b/counterexamples/phillips.lean @@ -283,7 +283,7 @@ begin by simp only [s, function.iterate_succ', union_comm, union_diff_self, subtype.coe_mk, union_diff_left], rw [nat.succ_eq_add_one, this, f.additive], - swap, { rw disjoint.comm, apply disjoint_diff }, + swap, { exact disjoint_sdiff_self_left }, calc ((n + 1 : ℕ) : ℝ) * (ε / 2) = ε / 2 + n * (ε / 2) : by simp only [nat.cast_succ]; ring ... ≤ f (↑(s (n + 1 : ℕ)) \ ↑(s n)) + f (↑(s n)) : add_le_add (I1 n) IH } }, @@ -339,7 +339,7 @@ begin simp only [discrete_part, continuous_part, restrict_apply], rw [← f.additive, ← inter_distrib_right], { simp only [union_univ, union_diff_self, univ_inter] }, - { have : disjoint f.discrete_support (univ \ f.discrete_support) := disjoint_diff, + { have : disjoint f.discrete_support (univ \ f.discrete_support) := disjoint_sdiff_self_right, exact this.mono (inter_subset_left _ _) (inter_subset_left _ _) } end @@ -361,7 +361,7 @@ begin conv_rhs { rw ← diff_union_inter t s }, rw [additive, self_eq_add_right], { exact continuous_part_apply_eq_zero_of_countable _ _ (hs.mono (inter_subset_right _ _)) }, - { exact disjoint.mono_right (inter_subset_right _ _) (disjoint.comm.1 disjoint_diff) }, + { exact disjoint.mono_right (inter_subset_right _ _) disjoint_sdiff_self_left }, end end bounded_additive_measure @@ -515,7 +515,7 @@ begin have : φ (f Hcont x) = ψ (spf Hcont x) := rfl, have U : univ = spf Hcont x ∪ (univ \ spf Hcont x), by simp only [union_univ, union_diff_self], rw [this, eq_add_parts, discrete_part_apply, hx, ψ.empty, zero_add, U, - ψ.continuous_part.additive _ _ (disjoint_diff), + ψ.continuous_part.additive _ _ disjoint_sdiff_self_right, ψ.continuous_part_apply_eq_zero_of_countable _ (countable_compl_spf Hcont x), add_zero], end diff --git a/src/algebra/indicator_function.lean b/src/algebra/indicator_function.lean index c28ce86dcc9d0..cc4c221a06b5f 100644 --- a/src/algebra/indicator_function.lean +++ b/src/algebra/indicator_function.lean @@ -364,8 +364,8 @@ by rw [sub_eq_add_neg, indicator_compl'] @[to_additive indicator_diff'] lemma mul_indicator_diff (h : s ⊆ t) (f : α → G) : mul_indicator (t \ s) f = mul_indicator t f * (mul_indicator s f)⁻¹ := -eq_mul_inv_of_mul_eq $ by rw [pi.mul_def, ← mul_indicator_union_of_disjoint disjoint_diff.symm f, - diff_union_self, union_eq_self_of_subset_right h] +eq_mul_inv_of_mul_eq $ by { rw [pi.mul_def, ←mul_indicator_union_of_disjoint, diff_union_self, + union_eq_self_of_subset_right h], exact disjoint_sdiff_self_left } lemma indicator_diff {G : Type*} [add_group G] {s t : set α} (h : s ⊆ t) (f : α → G) : indicator (t \ s) f = indicator t f - indicator s f := diff --git a/src/analysis/box_integral/partition/basic.lean b/src/analysis/box_integral/partition/basic.lean index 1dbd30f876aa9..e9ff3fb379ac2 100644 --- a/src/analysis/box_integral/partition/basic.lean +++ b/src/analysis/box_integral/partition/basic.lean @@ -631,7 +631,7 @@ lemma Union_bUnion_partition (h : ∀ J ∈ π, (πi J).is_partition) : (π.bUni Union_congr_of_surjective id surjective_id $ λ hJ, (h J hJ).Union_eq lemma is_partition_disj_union_of_eq_diff (h : π₂.Union = I \ π₁.Union) : - is_partition (π₁.disj_union π₂ (h.symm ▸ disjoint_diff)) := + is_partition (π₁.disj_union π₂ $ h.symm ▸ disjoint_sdiff_self_right) := is_partition_iff_Union_eq.2 $ (Union_disj_union _).trans $ by simp [h, π₁.Union_subset] end prepartition diff --git a/src/analysis/box_integral/partition/subbox_induction.lean b/src/analysis/box_integral/partition/subbox_induction.lean index 739d1b473f00f..e171fc752b88e 100644 --- a/src/analysis/box_integral/partition/subbox_induction.lean +++ b/src/analysis/box_integral/partition/subbox_induction.lean @@ -205,7 +205,7 @@ def union_compl_to_subordinate (π₁ : tagged_prepartition I) (π₂ : preparti (hU : π₂.Union = I \ π₁.Union) (r : (ι → ℝ) → Ioi (0 : ℝ)) : tagged_prepartition I := π₁.disj_union (π₂.to_subordinate r) - (((π₂.Union_to_subordinate r).trans hU).symm ▸ disjoint_diff) + (((π₂.Union_to_subordinate r).trans hU).symm ▸ disjoint_sdiff_self_right) lemma is_partition_union_compl_to_subordinate (π₁ : tagged_prepartition I) (π₂ : prepartition I) (hU : π₂.Union = I \ π₁.Union) (r : (ι → ℝ) → Ioi (0 : ℝ)) : diff --git a/src/data/finset/basic.lean b/src/data/finset/basic.lean index 724ecea454c5e..cb20eea268afe 100644 --- a/src/data/finset/basic.lean +++ b/src/data/finset/basic.lean @@ -276,8 +276,7 @@ theorem subset_iff {s₁ s₂ : finset α} : s₁ ⊆ s₂ ↔ ∀ ⦃x⦄, x theorem subset.antisymm_iff {s₁ s₂ : finset α} : s₁ = s₂ ↔ s₁ ⊆ s₂ ∧ s₂ ⊆ s₁ := le_antisymm_iff -theorem not_subset (s t : finset α) : ¬(s ⊆ t) ↔ ∃ x ∈ s, ¬(x ∈ t) := -by simp only [←finset.coe_subset, set.not_subset, exists_prop, finset.mem_coe] +lemma not_subset : ¬ s ⊆ t ↔ ∃ x ∈ s, x ∉ t := by simp only [←coe_subset, set.not_subset, mem_coe] @[simp] theorem le_eq_subset : ((≤) : finset α → finset α → Prop) = (⊆) := rfl @[simp] theorem lt_eq_subset : ((<) : finset α → finset α → Prop) = (⊂) := rfl @@ -438,7 +437,7 @@ end empty /-! ### singleton -/ section singleton -variables {s : finset α} {a : α} +variables {s : finset α} {a b : α} /-- `{a} : finset a` is the set `{a}` containing `a` and nothing else. @@ -460,8 +459,7 @@ theorem mem_singleton_self (a : α) : a ∈ ({a} : finset α) := or.inl rfl lemma singleton_injective : injective (singleton : α → finset α) := λ a b h, mem_singleton.1 (h ▸ mem_singleton_self _) -theorem singleton_inj {a b : α} : ({a} : finset α) = {b} ↔ a = b := -singleton_injective.eq_iff +@[simp] lemma singleton_inj : ({a} : finset α) = {b} ↔ a = b := singleton_injective.eq_iff @[simp] theorem singleton_nonempty (a : α) : ({a} : finset α).nonempty := ⟨a, mem_singleton_self a⟩ @@ -512,6 +510,8 @@ singleton_subset_set_iff @[simp] lemma subset_singleton_iff {s : finset α} {a : α} : s ⊆ {a} ↔ s = ∅ ∨ s = {a} := by rw [←coe_subset, coe_singleton, set.subset_singleton_iff_eq, coe_eq_empty, coe_eq_singleton] +lemma singleton_subset_singleton : ({a} : finset α) ⊆ {b} ↔ a = b := by simp + protected lemma nonempty.subset_singleton_iff {s : finset α} {a : α} (h : s.nonempty) : s ⊆ {a} ↔ s = {a} := subset_singleton_iff.trans $ or_iff_right h.ne_empty @@ -529,6 +529,10 @@ ssubset_singleton_iff.1 hs lemma eq_singleton_or_nontrivial (ha : a ∈ s) : s = {a} ∨ (s : set α).nontrivial := by { rw ←coe_eq_singleton, exact set.eq_singleton_or_nontrivial ha } +lemma nonempty.exists_eq_singleton_or_nontrivial : + s.nonempty → (∃ a, s = {a}) ∨ (s : set α).nontrivial := +λ ⟨a, ha⟩, (eq_singleton_or_nontrivial ha).imp_left $ exists.intro a + instance [nonempty α] : nontrivial (finset α) := ‹nonempty α›.elim $ λ a, ⟨⟨{a}, ∅, singleton_ne_empty _⟩⟩ @@ -576,7 +580,7 @@ by rwa [← coe_subset, coe_cons, coe_cons, set.insert_subset_insert_iff, coe_su lemma ssubset_iff_exists_cons_subset : s ⊂ t ↔ ∃ a (h : a ∉ s), s.cons a h ⊆ t := begin refine ⟨λ h, _, λ ⟨a, ha, h⟩, ssubset_of_ssubset_of_subset (ssubset_cons _) h⟩, - obtain ⟨a, hs, ht⟩ := (not_subset _ _).1 h.2, + obtain ⟨a, hs, ht⟩ := not_subset.1 h.2, exact ⟨a, ht, cons_subset.2 ⟨hs, h.subset⟩⟩, end @@ -592,12 +596,12 @@ lemma disjoint_left : disjoint s t ↔ ∀ ⦃a⦄, a ∈ s → a ∉ t := singleton_subset_iff.mp (h (singleton_subset_iff.mpr hs) (singleton_subset_iff.mpr ht)), λ h x hs ht a ha, h (hs ha) (ht ha)⟩ -lemma disjoint_val : disjoint s t ↔ s.1.disjoint t.1 := disjoint_left - lemma disjoint_right : disjoint s t ↔ ∀ ⦃a⦄, a ∈ t → a ∉ s := by rw [disjoint.comm, disjoint_left] lemma disjoint_iff_ne : disjoint s t ↔ ∀ a ∈ s, ∀ b ∈ t, a ≠ b := by simp only [disjoint_left, imp_not_comm, forall_eq'] +@[simp] lemma disjoint_val : s.1.disjoint t.1 ↔ disjoint s t := disjoint_left.symm + lemma _root_.disjoint.forall_ne_finset (h : disjoint s t) (ha : a ∈ s) (hb : b ∈ t) : a ≠ b := disjoint_iff_ne.1 h _ ha _ hb @@ -639,7 +643,7 @@ end disjoint It is the same as `s ∪ t`, but it does not require decidable equality on the type. The hypothesis ensures that the sets are disjoint. -/ def disj_union (s t : finset α) (h : disjoint s t) : finset α := -⟨s.1 + t.1, multiset.nodup_add.2 ⟨s.2, t.2, disjoint_val.1 h⟩⟩ +⟨s.1 + t.1, multiset.nodup_add.2 ⟨s.2, t.2, disjoint_val.2 h⟩⟩ @[simp] theorem mem_disj_union {α s t h a} : a ∈ @disj_union α s t h ↔ a ∈ s ∨ a ∈ t := @@ -1113,6 +1117,8 @@ end lemma inter_subset_inter_left (h : t ⊆ u) : s ∩ t ⊆ s ∩ u := inter_subset_inter subset.rfl h lemma inter_subset_inter_right (h : s ⊆ t) : s ∩ u ⊆ t ∩ u := inter_subset_inter h subset.rfl +lemma inter_subset_union : s ∩ t ⊆ s ∪ t := le_iff_subset.1 inf_le_sup + instance : distrib_lattice (finset α) := { le_sup_inf := assume a b c, show (a ∪ b) ∩ (a ∪ c) ⊆ a ∪ b ∩ c, by simp only [subset_iff, mem_inter, mem_union, and_imp, or_imp_distrib] {contextual:=tt}; @@ -1157,8 +1163,8 @@ lemma union_eq_empty_iff (A B : finset α) : A ∪ B = ∅ ↔ A = ∅ ∧ B = lemma union_subset_iff : s ∪ t ⊆ u ↔ s ⊆ u ∧ t ⊆ u := (sup_le_iff : s ⊔ t ≤ u ↔ s ≤ u ∧ t ≤ u) lemma subset_inter_iff : s ⊆ t ∩ u ↔ s ⊆ t ∧ s ⊆ u := (le_inf_iff : s ≤ t ⊓ u ↔ s ≤ t ∧ s ≤ u) -lemma inter_eq_left_iff_subset (s t : finset α) : s ∩ t = s ↔ s ⊆ t := inf_eq_left -lemma inter_eq_right_iff_subset (s t : finset α) : t ∩ s = s ↔ s ⊆ t := inf_eq_right +@[simp] lemma inter_eq_left_iff_subset (s t : finset α) : s ∩ t = s ↔ s ⊆ t := inf_eq_left +@[simp] lemma inter_eq_right_iff_subset (s t : finset α) : t ∩ s = s ↔ s ⊆ t := inf_eq_right lemma inter_congr_left (ht : s ∩ u ⊆ t) (hu : s ∩ t ⊆ u) : s ∩ t = s ∩ u := inf_congr_left ht hu lemma inter_congr_right (hs : t ∩ u ⊆ s) (ht : s ∩ u ⊆ t) : s ∩ u = t ∩ u := inf_congr_right hs ht @@ -1172,6 +1178,14 @@ lemma ite_subset_union (s s' : finset α) (P : Prop) [decidable P] : lemma inter_subset_ite (s s' : finset α) (P : Prop) [decidable P] : s ∩ s' ⊆ ite P s s' := inf_le_ite s s' P +lemma not_disjoint_iff_nonempty_inter : ¬disjoint s t ↔ (s ∩ t).nonempty := +not_disjoint_iff.trans $ by simp [finset.nonempty] + +alias not_disjoint_iff_nonempty_inter ↔ _ nonempty.not_disjoint + +lemma disjoint_or_nonempty_inter (s t : finset α) : disjoint s t ∨ (s ∩ t).nonempty := +by { rw ←not_disjoint_iff_nonempty_inter, exact em _ } + end lattice /-! ### erase -/ @@ -1256,7 +1270,7 @@ calc s.erase a ⊂ insert a (s.erase a) : ssubset_insert $ not_mem_erase _ _ lemma ssubset_iff_exists_subset_erase {s t : finset α} : s ⊂ t ↔ ∃ a ∈ t, s ⊆ t.erase a := begin refine ⟨λ h, _, λ ⟨a, ha, h⟩, ssubset_of_subset_of_ssubset h $ erase_ssubset ha⟩, - obtain ⟨a, ht, hs⟩ := (not_subset _ _).1 h.2, + obtain ⟨a, ht, hs⟩ := not_subset.1 h.2, exact ⟨a, ht, subset_erase.2 ⟨h.1, hs⟩⟩, end @@ -1374,7 +1388,9 @@ lemma sdiff_union_inter (s t : finset α) : (s \ t) ∪ (s ∩ t) = s := sup_sdi @[simp] lemma sdiff_idem (s t : finset α) : s \ t \ t = s \ t := sdiff_idem -lemma sdiff_eq_empty_iff_subset : s \ t = ∅ ↔ s ⊆ t := sdiff_eq_bot_iff +lemma subset_sdiff : s ⊆ t \ u ↔ s ⊆ t ∧ disjoint s u := le_iff_subset.symm.trans le_sdiff + +@[simp] lemma sdiff_eq_empty_iff_subset : s \ t = ∅ ↔ s ⊆ t := sdiff_eq_bot_iff lemma sdiff_nonempty : (s \ t).nonempty ↔ ¬ s ⊆ t := nonempty_iff_ne_empty.trans sdiff_eq_empty_iff_subset.not @@ -2218,7 +2234,7 @@ hypothesis ensures that the sets are disjoint. -/ def disj_Union (s : finset α) (t : α → finset β) (hf : (s : set α).pairwise_disjoint t) : finset β := ⟨(s.val.bind (finset.val ∘ t)), multiset.nodup_bind.mpr - ⟨λ a ha, (t a).nodup, s.nodup.pairwise $ λ a ha b hb hab, finset.disjoint_val.1 $ hf ha hb hab⟩⟩ + ⟨λ a ha, (t a).nodup, s.nodup.pairwise $ λ a ha b hb hab, disjoint_val.2 $ hf ha hb hab⟩⟩ @[simp] theorem disj_Union_val (s : finset α) (t : α → finset β) (h) : (s.disj_Union t h).1 = (s.1.bind (λ a, (t a).1)) := rfl diff --git a/src/data/finset/card.lean b/src/data/finset/card.lean index c608b4cd382bf..8a6f1f2e0e381 100644 --- a/src/data/finset/card.lean +++ b/src/data/finset/card.lean @@ -184,6 +184,12 @@ card_le_of_subset $ filter_subset _ _ lemma eq_of_subset_of_card_le {s t : finset α} (h : s ⊆ t) (h₂ : t.card ≤ s.card) : s = t := eq_of_veq $ multiset.eq_of_le_of_card_le (val_le_iff.mpr h) h₂ +lemma eq_of_superset_of_card_ge (hst : s ⊆ t) (hts : t.card ≤ s.card) : t = s := +(eq_of_subset_of_card_le hst hts).symm + +lemma subset_iff_eq_of_card_le (h : t.card ≤ s.card) : s ⊆ t ↔ s = t := +⟨λ hst, eq_of_subset_of_card_le hst h, eq.subset'⟩ + lemma map_eq_of_subset {f : α ↪ α} (hs : s.map f ⊆ s) : s.map f = s := eq_of_subset_of_card_le hs (card_map _).ge @@ -301,6 +307,9 @@ variables [decidable_eq α] lemma card_union_add_card_inter (s t : finset α) : (s ∪ t).card + (s ∩ t).card = s.card + t.card := finset.induction_on t (by simp) $ λ a r har, by by_cases a ∈ s; simp *; cc +lemma card_inter_add_card_union (s t : finset α) : (s ∩ t).card + (s ∪ t).card = s.card + t.card := +by rw [add_comm, card_union_add_card_inter] + lemma card_union_le (s t : finset α) : (s ∪ t).card ≤ s.card + t.card := card_union_add_card_inter s t ▸ nat.le_add_right _ _ @@ -323,6 +332,9 @@ calc card t - card s ... = card (t \ (s ∩ t)) : (card_sdiff (inter_subset_right s t)).symm ... ≤ card (t \ s) : by rw sdiff_inter_self_right t s +lemma card_le_card_sdiff_add_card : s.card ≤ (s \ t).card + t.card := +tsub_le_iff_right.1 $ le_card_sdiff _ _ + lemma card_sdiff_add_card : (s \ t).card + t.card = (s ∪ t).card := by rw [←card_disjoint_union sdiff_disjoint, sdiff_union_self_eq_union] diff --git a/src/data/finset/image.lean b/src/data/finset/image.lean index 3696de6c34e8c..601b3bbd93e24 100644 --- a/src/data/finset/image.lean +++ b/src/data/finset/image.lean @@ -239,6 +239,9 @@ by simp only [mem_def, image_val, mem_dedup, multiset.mem_map, exists_prop] lemma mem_image_of_mem (f : α → β) {a} (h : a ∈ s) : f a ∈ s.image f := mem_image.2 ⟨_, h, rfl⟩ +lemma forall_image {p : β → Prop} : (∀ b ∈ s.image f, p b) ↔ ∀ a ∈ s, p (f a) := +by simp only [mem_image, forall_exists_index, forall_apply_eq_imp_iff₂] + @[simp] lemma mem_image_const : c ∈ s.image (const α b) ↔ s.nonempty ∧ b = c := by { rw mem_image, simp only [exists_prop, const_apply, exists_and_distrib_right], refl } diff --git a/src/data/finset/lattice.lean b/src/data/finset/lattice.lean index 8ba891c9b7c67..321f110d5ab60 100644 --- a/src/data/finset/lattice.lean +++ b/src/data/finset/lattice.lean @@ -26,7 +26,7 @@ variables [semilattice_sup α] [order_bot α] /-- Supremum of a finite set: `sup {a, b, c} f = f a ⊔ f b ⊔ f c` -/ def sup (s : finset β) (f : β → α) : α := s.fold (⊔) ⊥ f -variables {s s₁ s₂ : finset β} {f g : β → α} +variables {s s₁ s₂ : finset β} {f g : β → α} {a : α} lemma sup_def : s.sup f = (s.1.map f).sup := rfl @@ -76,6 +76,8 @@ alias finset.sup_le_iff ↔ _ sup_le attribute [protected] sup_le +lemma sup_const_le : s.sup (λ _, a) ≤ a := finset.sup_le $ λ _ _, le_rfl + lemma le_sup {b : β} (hb : b ∈ s) : f b ≤ s.sup f := finset.sup_le_iff.1 le_rfl _ hb @[simp] lemma sup_bUnion [decidable_eq β] (s : finset γ) (t : γ → finset β) : @@ -252,7 +254,7 @@ variables [semilattice_inf α] [order_top α] /-- Infimum of a finite set: `inf {a, b, c} f = f a ⊓ f b ⊓ f c` -/ def inf (s : finset β) (f : β → α) : α := s.fold (⊓) ⊤ f -variables {s s₁ s₂ : finset β} {f g : β → α} +variables {s s₁ s₂ : finset β} {f g : β → α} {a : α} lemma inf_def : s.inf f = (s.1.map f).inf := rfl @@ -301,6 +303,8 @@ alias finset.le_inf_iff ↔ _ le_inf attribute [protected] le_inf +lemma le_inf_const_le : a ≤ s.inf (λ _, a) := finset.le_inf $ λ _ _, le_rfl + lemma inf_le {b : β} (hb : b ∈ s) : s.inf f ≤ f b := finset.le_inf_iff.1 le_rfl _ hb lemma inf_mono_fun {g : β → α} (h : ∀ b ∈ s, f b ≤ g b) : s.inf f ≤ s.inf g := diff --git a/src/data/fintype/basic.lean b/src/data/fintype/basic.lean index 582ec043777c0..3c71e4fdf1e89 100644 --- a/src/data/fintype/basic.lean +++ b/src/data/fintype/basic.lean @@ -52,7 +52,7 @@ class fintype (α : Type*) := (complete : ∀ x : α, x ∈ elems) namespace finset -variables [fintype α] {s : finset α} +variables [fintype α] {s t : finset α} /-- `univ` is the universal finite set of type `finset α` implied from the assumption `fintype α`. -/ @@ -98,6 +98,12 @@ instance : bounded_order (finset α) := lemma ssubset_univ_iff {s : finset α} : s ⊂ univ ↔ s ≠ univ := @lt_top_iff_ne_top _ _ _ s +lemma codisjoint_left : codisjoint s t ↔ ∀ ⦃a⦄, a ∉ s → a ∈ t := +by { classical, simp [codisjoint_iff, eq_univ_iff_forall, or_iff_not_imp_left] } + +lemma codisjoint_right : codisjoint s t ↔ ∀ ⦃a⦄, a ∉ t → a ∈ s := +codisjoint.comm.trans codisjoint_left + section boolean_algebra variables [decidable_eq α] {a : α} @@ -316,6 +322,16 @@ def of_surjective [decidable_eq β] [fintype α] (f : α → β) (H : function.s end fintype +namespace finset +variables [fintype α] [decidable_eq α] {s t : finset α} + +instance decidable_codisjoint : decidable (codisjoint s t) := +decidable_of_iff _ codisjoint_left.symm + +instance decidable_is_compl : decidable (is_compl s t) := decidable_of_iff' _ is_compl_iff + +end finset + section inv namespace function @@ -699,6 +715,9 @@ instance plift.fintype_Prop (p : Prop) [decidable p] : fintype (plift p) := instance Prop.fintype : fintype Prop := ⟨⟨{true, false}, by simp [true_ne_false]⟩, classical.cases (by simp) (by simp)⟩ +@[simp] lemma fintype.univ_Prop : (finset.univ : finset Prop) = {true, false} := +finset.eq_of_veq $ by simp; refl + instance subtype.fintype (p : α → Prop) [decidable_pred p] [fintype α] : fintype {x // p x} := fintype.subtype (univ.filter p) (by simp) diff --git a/src/data/set/basic.lean b/src/data/set/basic.lean index b31c965bf21e0..661e500e3b9cf 100644 --- a/src/data/set/basic.lean +++ b/src/data/set/basic.lean @@ -162,7 +162,7 @@ lemma eq.subset {α} {s t : set α} : s = t → s ⊆ t := eq.subset' namespace set -variables {α : Type u} {β : Type v} {γ : Type w} {ι : Sort x} {a b : α} {s t u : set α} +variables {α : Type u} {β : Type v} {γ : Type w} {ι : Sort x} {a b : α} {s s₁ s₂ t t₁ t₂ u : set α} instance : inhabited (set α) := ⟨∅⟩ @@ -914,6 +914,54 @@ lemma disjoint_left : disjoint s t ↔ ∀ ⦃a⦄, a ∈ s → a ∉ t := disjoint_iff_inf_le.trans $ forall_congr $ λ _, not_and lemma disjoint_right : disjoint s t ↔ ∀ ⦃a⦄, a ∈ t → a ∉ s := by rw [disjoint.comm, disjoint_left] +lemma not_disjoint_iff : ¬disjoint s t ↔ ∃ x, x ∈ s ∧ x ∈ t := +set.disjoint_iff.not.trans $ not_forall.trans $ exists_congr $ λ x, not_not + +lemma not_disjoint_iff_nonempty_inter : ¬disjoint s t ↔ (s ∩ t).nonempty := not_disjoint_iff + +alias not_disjoint_iff_nonempty_inter ↔ _ nonempty.not_disjoint + +lemma disjoint_or_nonempty_inter (s t : set α) : disjoint s t ∨ (s ∩ t).nonempty := +(em _).imp_right not_disjoint_iff_nonempty_inter.mp + +lemma disjoint_iff_forall_ne : disjoint s t ↔ ∀ (x ∈ s) (y ∈ t), x ≠ y := +by simp only [ne.def, disjoint_left, @imp_not_comm _ (_ = _), forall_eq'] + +lemma _root_.disjoint.ne_of_mem (h : disjoint s t) {x y} (hx : x ∈ s) (hy : y ∈ t) : x ≠ y := +disjoint_iff_forall_ne.mp h x hx y hy + +lemma disjoint_of_subset_left (hs : s₁ ⊆ s₂) (h : disjoint s₂ t) : disjoint s₁ t := h.mono_left hs +lemma disjoint_of_subset_right (ht : t₁ ⊆ t₂) (h : disjoint s t₂) : disjoint s t₁ := h.mono_right ht + +lemma disjoint_of_subset (hs : s₁ ⊆ s₂) (ht : t₁ ⊆ t₂) (h : disjoint s₂ t₂) : disjoint s₁ t₁ := +h.mono hs ht + +@[simp] lemma disjoint_union_left : disjoint (s ∪ t) u ↔ disjoint s u ∧ disjoint t u := +disjoint_sup_left + +@[simp] lemma disjoint_union_right : disjoint s (t ∪ u) ↔ disjoint s t ∧ disjoint s u := +disjoint_sup_right + +@[simp] lemma disjoint_empty (s : set α) : disjoint s ∅ := disjoint_bot_right +@[simp] lemma empty_disjoint (s : set α) : disjoint ∅ s := disjoint_bot_left + +@[simp] lemma univ_disjoint : disjoint univ s ↔ s = ∅ := top_disjoint +@[simp] lemma disjoint_univ : disjoint s univ ↔ s = ∅ := disjoint_top + +lemma disjoint_sdiff_left : disjoint (t \ s) s := disjoint_sdiff_self_left +lemma disjoint_sdiff_right : disjoint s (t \ s) := disjoint_sdiff_self_right + +@[simp] lemma disjoint_singleton_left : disjoint {a} s ↔ a ∉ s := +by simp [set.disjoint_iff, subset_def]; exact iff.rfl + +@[simp] lemma disjoint_singleton_right : disjoint s {a} ↔ a ∉ s := +disjoint.comm.trans disjoint_singleton_left + +@[simp] lemma disjoint_singleton : disjoint ({a} : set α) {b} ↔ a ≠ b := +by rw [disjoint_singleton_left, mem_singleton_iff] + +lemma subset_diff : s ⊆ t \ u ↔ s ⊆ t ∧ disjoint s u := le_iff_subset.symm.trans le_sdiff + /-! ### Lemmas about complement -/ lemma compl_def (s : set α) : sᶜ = {x | x ∉ s} := rfl @@ -1173,11 +1221,8 @@ sdiff_inf_self_right _ _ @[simp] theorem diff_self_inter {s t : set α} : s \ (s ∩ t) = s \ t := sdiff_inf_self_left _ _ -@[simp] theorem diff_eq_self {s t : set α} : s \ t = s ↔ t ∩ s ⊆ ∅ := -show s \ t = s ↔ t ⊓ s ≤ ⊥, from sdiff_eq_self_iff_disjoint.trans disjoint_iff_inf_le - @[simp] theorem diff_singleton_eq_self {a : α} {s : set α} (h : a ∉ s) : s \ {a} = s := -diff_eq_self.2 $ by simp [singleton_inter_eq_empty.2 h] +sdiff_eq_self_iff_disjoint.2 $ by simp [h] @[simp] theorem insert_diff_singleton {a : α} {s : set α} : insert a (s \ {a}) = insert a s := @@ -1558,6 +1603,12 @@ by simp [or_iff_not_imp_right] lemma eq_singleton_or_nontrivial (ha : a ∈ s) : s = {a} ∨ s.nontrivial := by { rw ←subsingleton_iff_singleton ha, exact s.subsingleton_or_nontrivial } +lemma nontrivial_iff_ne_singleton (ha : a ∈ s) : s.nontrivial ↔ s ≠ {a} := +⟨nontrivial.ne_singleton, (eq_singleton_or_nontrivial ha).resolve_left⟩ + +lemma nonempty.exists_eq_singleton_or_nontrivial : s.nonempty → (∃ a, s = {a}) ∨ s.nontrivial := +λ ⟨a, ha⟩, (eq_singleton_or_nontrivial ha).imp_left $ exists.intro a + theorem univ_eq_true_false : univ = ({true, false} : set Prop) := eq.symm $ eq_univ_of_forall $ classical.cases (by simp) (by simp) @@ -1793,8 +1844,6 @@ end monotone /-! ### Disjoint sets -/ -section disjoint - variables {α β : Type*} {s t u : set α} {f : α → β} namespace disjoint @@ -1824,70 +1873,3 @@ lemma subset_right_of_subset_union (h : s ⊆ t ∪ u) (hab : disjoint s t) : s hab.left_le_of_le_sup_left h end disjoint - -namespace set - -lemma not_disjoint_iff : ¬disjoint s t ↔ ∃ x, x ∈ s ∧ x ∈ t := -set.disjoint_iff.not.trans $ not_forall.trans $ exists_congr $ λ x, not_not - -lemma not_disjoint_iff_nonempty_inter : ¬disjoint s t ↔ (s ∩ t).nonempty := -not_disjoint_iff - -alias not_disjoint_iff_nonempty_inter ↔ _ nonempty.not_disjoint - -lemma disjoint_or_nonempty_inter (s t : set α) : disjoint s t ∨ (s ∩ t).nonempty := -(em _).imp_right not_disjoint_iff_nonempty_inter.mp - -lemma disjoint_iff_forall_ne : disjoint s t ↔ ∀ (x ∈ s) (y ∈ t), x ≠ y := -by simp only [ne.def, disjoint_left, @imp_not_comm _ (_ = _), forall_eq'] - -lemma _root_.disjoint.ne_of_mem (h : disjoint s t) {x y} (hx : x ∈ s) (hy : y ∈ t) : x ≠ y := -disjoint_iff_forall_ne.mp h x hx y hy - -theorem disjoint_of_subset_left (h : s ⊆ u) (d : disjoint u t) : disjoint s t := -d.mono_left h - -theorem disjoint_of_subset_right (h : t ⊆ u) (d : disjoint s u) : disjoint s t := -d.mono_right h - -theorem disjoint_of_subset {s t u v : set α} (h1 : s ⊆ u) (h2 : t ⊆ v) (d : disjoint u v) : - disjoint s t := -d.mono h1 h2 - -@[simp] theorem disjoint_union_left : - disjoint (s ∪ t) u ↔ disjoint s u ∧ disjoint t u := -disjoint_sup_left - -@[simp] theorem disjoint_union_right : - disjoint s (t ∪ u) ↔ disjoint s t ∧ disjoint s u := -disjoint_sup_right - -theorem disjoint_diff {a b : set α} : disjoint a (b \ a) := -disjoint_iff.2 (inter_diff_self _ _) - -@[simp] theorem disjoint_empty (s : set α) : disjoint s ∅ := disjoint_bot_right - -@[simp] theorem empty_disjoint (s : set α) : disjoint ∅ s := disjoint_bot_left - -@[simp] lemma univ_disjoint {s : set α} : disjoint univ s ↔ s = ∅ := -top_disjoint - -@[simp] lemma disjoint_univ {s : set α} : disjoint s univ ↔ s = ∅ := -disjoint_top - -@[simp] theorem disjoint_singleton_left {a : α} {s : set α} : disjoint {a} s ↔ a ∉ s := -by simp [set.disjoint_iff, subset_def]; exact iff.rfl - -@[simp] theorem disjoint_singleton_right {a : α} {s : set α} : disjoint s {a} ↔ a ∉ s := -by rw [disjoint.comm]; exact disjoint_singleton_left - -@[simp] lemma disjoint_singleton {a b : α} : disjoint ({a} : set α) {b} ↔ a ≠ b := -by rw [disjoint_singleton_left, mem_singleton_iff] - -lemma subset_diff {s t u : set α} : s ⊆ t \ u ↔ s ⊆ t ∧ disjoint s u := -⟨λ h, ⟨λ x hxs, (h hxs).1, disjoint_iff_inf_le.mpr $ λ x ⟨hxs, hxu⟩, (h hxs).2 hxu⟩, -λ ⟨h1, h2⟩ x hxs, ⟨h1 hxs, λ hxu, h2.le_bot ⟨hxs, hxu⟩⟩⟩ - -end set - -end disjoint diff --git a/src/data/set/lattice.lean b/src/data/set/lattice.lean index 9f70b6e4d79a0..d29f105c83895 100644 --- a/src/data/set/lattice.lean +++ b/src/data/set/lattice.lean @@ -312,9 +312,27 @@ lemma Inter_congr_of_surjective {f : ι → set α} {g : ι₂ → set α} (h : (h1 : surjective h) (h2 : ∀ x, g (h x) = f x) : (⋂ x, f x) = ⋂ y, g y := h1.infi_congr h h2 -theorem Union_const [nonempty ι] (s : set β) : (⋃ i : ι, s) = s := supr_const +lemma Union_congr {s t : ι → set α} (h : ∀ i, s i = t i) : (⋃ i, s i) = ⋃ i, t i := supr_congr h +lemma Inter_congr {s t : ι → set α} (h : ∀ i, s i = t i) : (⋂ i, s i) = ⋂ i, t i := infi_congr h + +lemma Union₂_congr {s t : Π i, κ i → set α} (h : ∀ i j, s i j = t i j) : + (⋃ i j, s i j) = ⋃ i j, t i j := +Union_congr $ λ i, Union_congr $ h i + +lemma Inter₂_congr {s t : Π i, κ i → set α} (h : ∀ i j, s i j = t i j) : + (⋂ i j, s i j) = ⋂ i j, t i j := +Inter_congr $ λ i, Inter_congr $ h i + +section nonempty +variables [nonempty ι] {f : ι → set α} {s : set α} + +lemma Union_const (s : set β) : (⋃ i : ι, s) = s := supr_const +lemma Inter_const (s : set β) : (⋂ i : ι, s) = s := infi_const -theorem Inter_const [nonempty ι] (s : set β) : (⋂ i : ι, s) = s := infi_const +lemma Union_eq_const (hf : ∀ i, f i = s) : (⋃ i, f i) = s := (Union_congr hf).trans $ Union_const _ +lemma Inter_eq_const (hf : ∀ i, f i = s) : (⋂ i, f i) = s := (Inter_congr hf).trans $ Inter_const _ + +end nonempty @[simp] theorem compl_Union (s : ι → set β) : (⋃ i, s i)ᶜ = (⋂ i, (s i)ᶜ) := compl_supr @@ -614,17 +632,6 @@ lemma bInter_mono {s s' : set α} {t t' : α → set β} (hs : s ⊆ s') (h : (⋂ x ∈ s', t x) ⊆ (⋂ x ∈ s, t' x) := (bInter_subset_bInter_left hs).trans $ Inter₂_mono h -lemma Union_congr {s t : ι → set α} (h : ∀ i, s i = t i) : (⋃ i, s i) = ⋃ i, t i := supr_congr h -lemma Inter_congr {s t : ι → set α} (h : ∀ i, s i = t i) : (⋂ i, s i) = ⋂ i, t i := infi_congr h - -lemma Union₂_congr {s t : Π i, κ i → set α} (h : ∀ i j, s i j = t i j) : - (⋃ i j, s i j) = ⋃ i j, t i j := -Union_congr $ λ i, Union_congr $ h i - -lemma Inter₂_congr {s t : Π i, κ i → set α} (h : ∀ i j, s i j = t i j) : - (⋂ i j, s i j) = ⋂ i j, t i j := -Inter_congr $ λ i, Inter_congr $ h i - theorem bUnion_eq_Union (s : set α) (t : Π x ∈ s, set β) : (⋃ x ∈ s, t x ‹_›) = (⋃ x : s, t x x.2) := supr_subtype' diff --git a/src/group_theory/perm/cycle/type.lean b/src/group_theory/perm/cycle/type.lean index 3b66c3bc4cd65..9069d0086b584 100644 --- a/src/group_theory/perm/cycle/type.lean +++ b/src/group_theory/perm/cycle/type.lean @@ -112,8 +112,7 @@ lemma disjoint.cycle_type {σ τ : perm α} (h : disjoint σ τ) : begin rw [cycle_type_def, cycle_type_def, cycle_type_def, h.cycle_factors_finset_mul_eq_union, ←multiset.map_add, finset.union_val, multiset.add_eq_union_iff_disjoint.mpr _], - rw [←finset.disjoint_val], - exact h.disjoint_cycle_factors_finset + exact finset.disjoint_val.2 h.disjoint_cycle_factors_finset end lemma cycle_type_inv (σ : perm α) : σ⁻¹.cycle_type = σ.cycle_type := diff --git a/src/measure_theory/decomposition/jordan.lean b/src/measure_theory/decomposition/jordan.lean index 81bbab97509d1..a332a7ca18a72 100644 --- a/src/measure_theory/decomposition/jordan.lean +++ b/src/measure_theory/decomposition/jordan.lean @@ -234,7 +234,7 @@ lemma subset_positive_null_set (hsu : 0 ≤[u] s) (hw₁ : s w = 0) (hw₂ : w ⊆ u) (hwt : v ⊆ w) : s v = 0 := begin have : s v + s (w \ v) = 0, - { rw [← hw₁, ← of_union set.disjoint_diff hv (hw.diff hv), + { rw [← hw₁, ← of_union set.disjoint_sdiff_right hv (hw.diff hv), set.union_diff_self, set.union_eq_self_of_subset_left hwt], apply_instance }, have h₁ := nonneg_of_zero_le_restrict _ (restrict_le_restrict_subset _ _ hu hsu (hwt.trans hw₂)), @@ -264,7 +264,7 @@ begin rw restrict_le_restrict_iff at hsu hsv, have a := hsu (hu.diff hv) (u.diff_subset v), have b := hsv (hv.diff hu) (v.diff_subset u), - erw [of_union (set.disjoint_of_subset_left (u.diff_subset v) set.disjoint_diff) + erw [of_union (set.disjoint_of_subset_left (u.diff_subset v) disjoint_sdiff_self_right) (hu.diff hv) (hv.diff hu)] at hs, rw zero_apply at a b, split, diff --git a/src/measure_theory/decomposition/signed_hahn.lean b/src/measure_theory/decomposition/signed_hahn.lean index af0b358545896..25a9760d1e745 100644 --- a/src/measure_theory/decomposition/signed_hahn.lean +++ b/src/measure_theory/decomposition/signed_hahn.lean @@ -369,7 +369,7 @@ begin { intro n, refine le_trans _ (le_of_lt (h_lt _)), rw [hA, ← set.diff_union_of_subset (set.subset_Union _ n), - of_union (disjoint.comm.1 set.disjoint_diff) _ (hmeas n)], + of_union set.disjoint_sdiff_left _ (hmeas n)], { refine add_le_of_nonpos_left _, have : s ≤[A] 0 := restrict_le_restrict_Union _ _ hmeas hr, refine nonpos_of_restrict_le_zero _ (restrict_le_zero_subset _ _ (set.diff_subset _ _) this), @@ -397,7 +397,7 @@ begin { apply le_antisymm, { refine le_of_tendsto_of_tendsto tendsto_const_nhds hf₂ (eventually_of_forall (λ n, _)), rw [← (hB n).2, hA, ← set.diff_union_of_subset (set.subset_Union _ n), - of_union (disjoint.comm.1 set.disjoint_diff) _ (hB₁ n)], + of_union set.disjoint_sdiff_left _ (hB₁ n)], { refine add_le_of_nonpos_left _, have : s ≤[A] 0 := restrict_le_restrict_Union _ _ hB₁ (λ m, let ⟨_, h⟩ := (hB m).1 in h), diff --git a/src/measure_theory/decomposition/unsigned_hahn.lean b/src/measure_theory/decomposition/unsigned_hahn.lean index 3dc48900f73ed..e2d80b9594cf2 100644 --- a/src/measure_theory/decomposition/unsigned_hahn.lean +++ b/src/measure_theory/decomposition/unsigned_hahn.lean @@ -190,11 +190,8 @@ begin exact ((add_le_add_iff_left γ).1 $ calc γ + d t ≤ d s + d t : add_le_add γ_le_d_s le_rfl ... = d (s ∪ t) : - begin - rw [d_split _ _ (hs.union ht) ht, union_diff_right, union_inter_cancel_right, - diff_eq_self.2], - exact assume a ⟨hat, has⟩, hts hat has - end + by rw [d_split _ _ (hs.union ht) ht, union_diff_right, union_inter_cancel_right, + (subset_compl_iff_disjoint_left.1 hts).sdiff_eq_left] ... ≤ γ + 0 : by rw [add_zero]; exact d_le_γ _ (hs.union ht)), rw [← to_nnreal_μ, ← to_nnreal_ν, ennreal.coe_le_coe, ← nnreal.coe_le_coe], simpa only [d, sub_le_iff_le_add, zero_add] using this } diff --git a/src/measure_theory/integral/set_integral.lean b/src/measure_theory/integral/set_integral.lean index 67350e7fbc6f7..1ad13b8d3bbbc 100644 --- a/src/measure_theory/integral/set_integral.lean +++ b/src/measure_theory/integral/set_integral.lean @@ -88,7 +88,7 @@ lemma integral_diff (ht : measurable_set t) (hfs : integrable_on f s μ) ∫ x in s \ t, f x ∂μ = ∫ x in s, f x ∂μ - ∫ x in t, f x ∂μ := begin rw [eq_sub_iff_add_eq, ← integral_union, diff_union_of_subset hts], - exacts [disjoint_diff.symm, ht, hfs.mono_set (diff_subset _ _), hft] + exacts [disjoint_sdiff_self_left, ht, hfs.mono_set (diff_subset _ _), hft] end lemma integral_finset_bUnion {ι : Type*} (t : finset ι) {s : ι → set α} @@ -239,7 +239,7 @@ lemma set_integral_union_eq_left {f : α → E} (hf : strongly_measurable f) (hf begin rw [← set.union_diff_self, union_comm, integral_union, set_integral_eq_zero_of_forall_eq_zero _ (λ x hx, ht_eq x (diff_subset _ _ hx)), zero_add], - exacts [hf, disjoint_diff.symm, hs, hfi.integrable_on, hfi.integrable_on] + exacts [hf, disjoint_sdiff_self_left, hs, hfi.integrable_on, hfi.integrable_on] end lemma set_integral_neg_eq_set_integral_nonpos [linear_order E] [order_closed_topology E] diff --git a/src/measure_theory/measure/ae_disjoint.lean b/src/measure_theory/measure/ae_disjoint.lean index 50d95d8d859de..f9148196a88b3 100644 --- a/src/measure_theory/measure/ae_disjoint.lean +++ b/src/measure_theory/measure/ae_disjoint.lean @@ -110,7 +110,8 @@ set `u`. -/ lemma exists_disjoint_diff (h : ae_disjoint μ s t) : ∃ u, measurable_set u ∧ μ u = 0 ∧ disjoint (s \ u) t := ⟨to_measurable μ (s ∩ t), measurable_set_to_measurable _ _, (measure_to_measurable _).trans h, - disjoint_diff.symm.mono_left (λ x hx, ⟨hx.1, λ hxt, hx.2 $ subset_to_measurable _ _ ⟨hx.1, hxt⟩⟩)⟩ + disjoint_sdiff_self_left.mono_left $ λ x hx, ⟨hx.1, λ hxt, hx.2 $ + subset_to_measurable _ _ ⟨hx.1, hxt⟩⟩⟩ lemma of_null_right (h : μ t = 0) : ae_disjoint μ s t := measure_mono_null (inter_subset_right _ _) h diff --git a/src/measure_theory/measure/measure_space.lean b/src/measure_theory/measure/measure_space.lean index fdcf54b5b8a7f..75c82f98d98a1 100644 --- a/src/measure_theory/measure/measure_space.lean +++ b/src/measure_theory/measure/measure_space.lean @@ -212,7 +212,7 @@ lemma measure_diff_null (h : μ s₂ = 0) : μ (s₁ \ s₂) = μ s₁ := measure_diff_null' $ measure_mono_null (inter_subset_right _ _) h lemma measure_add_diff (hs : measurable_set s) (t : set α) : μ s + μ (t \ s) = μ (s ∪ t) := -by rw [← measure_union' disjoint_diff hs, union_diff_self] +by rw [← measure_union' disjoint_sdiff_right hs, union_diff_self] lemma measure_diff' (s : set α) (hm : measurable_set t) (h_fin : μ t ≠ ∞) : μ (s \ t) = μ (s ∪ t) - μ t := @@ -474,7 +474,7 @@ begin use j, rw [← measure_diff hjk (h _) (this _ hjk)], exact measure_mono (diff_subset_diff_right hji) }, - { rw [tsub_le_iff_right, ← measure_union disjoint_diff.symm (h i), set.union_comm], + { rw [tsub_le_iff_right, ← measure_union disjoint_sdiff_left (h i), set.union_comm], exact measure_mono (diff_subset_iff.1 $ subset.refl _) } }, { exact hd.mono_comp _ (λ _ _, diff_subset_diff_right) } end diff --git a/src/measure_theory/measure/outer_measure.lean b/src/measure_theory/measure/outer_measure.lean index 7b1077ebef4d0..1c03e3594c176 100644 --- a/src/measure_theory/measure/outer_measure.lean +++ b/src/measure_theory/measure/outer_measure.lean @@ -1254,7 +1254,7 @@ lemma extend_mono {s₁ s₂ : set α} (h₁ : measurable_set s₁) (hs : s₁ extend m s₁ ≤ extend m s₂ := begin refine le_infi _, intro h₂, - have := extend_union measurable_set.empty m0 measurable_set.Union mU disjoint_diff + have := extend_union measurable_set.empty m0 measurable_set.Union mU disjoint_sdiff_self_right h₁ (h₂.diff h₁), rw union_diff_cancel hs at this, rw ← extend_eq m, diff --git a/src/measure_theory/measure/vector_measure.lean b/src/measure_theory/measure/vector_measure.lean index 720678d4bc1d5..ea7e57ae0726c 100644 --- a/src/measure_theory/measure/vector_measure.lean +++ b/src/measure_theory/measure/vector_measure.lean @@ -174,7 +174,7 @@ end lemma of_add_of_diff {A B : set α} (hA : measurable_set A) (hB : measurable_set B) (h : A ⊆ B) : v A + v (B \ A) = v B := begin - rw [← of_union disjoint_diff hA (hB.diff hA), union_diff_cancel h], + rw [← of_union disjoint_sdiff_right hA (hB.diff hA), union_diff_cancel h], apply_instance, end @@ -196,12 +196,12 @@ begin ... = v (A \ B) + v (A ∩ B) : by { rw of_union, { rw disjoint.comm, - exact set.disjoint_of_subset_left (A.inter_subset_right B) set.disjoint_diff }, + exact set.disjoint_of_subset_left (A.inter_subset_right B) disjoint_sdiff_self_right }, { exact hA.diff hB }, { exact hA.inter hB } } ... = v (A \ B) + v (A ∩ B ∪ B \ A) : by { rw [of_union, h', add_zero], - { exact set.disjoint_of_subset_left (A.inter_subset_left B) set.disjoint_diff }, + { exact set.disjoint_of_subset_left (A.inter_subset_left B) disjoint_sdiff_self_right }, { exact hA.inter hB }, { exact hB.diff hA } } ... = v (A \ B) + v B : @@ -1144,7 +1144,7 @@ begin { exact subset.trans (inter_subset_left _ _) (diff_subset _ _) }, { exact inter_subset_left _ _ }, { apply_instance }, - { exact disjoint.mono (inter_subset_left _ _) (inter_subset_left _ _) disjoint_diff }, + { exact disjoint_sdiff_self_right.mono (inter_subset_left _ _) (inter_subset_left _ _) }, { apply subset.antisymm; intros x hx, { by_cases hxu' : x ∈ uᶜ, diff --git a/src/order/boolean_algebra.lean b/src/order/boolean_algebra.lean index f0c33b28e318f..b83a09204696b 100644 --- a/src/order/boolean_algebra.lean +++ b/src/order/boolean_algebra.lean @@ -176,6 +176,13 @@ disjoint_iff_inf_le.mpr inf_sdiff_self_left.le theorem disjoint_sdiff_self_right : disjoint x (y \ x) := disjoint_iff_inf_le.mpr inf_sdiff_self_right.le +lemma le_sdiff : x ≤ y \ z ↔ x ≤ y ∧ disjoint x z := +⟨λ h, ⟨h.trans sdiff_le, disjoint_sdiff_self_left.mono_left h⟩, λ h, + by { rw ←h.2.sdiff_eq_left, exact sdiff_le_sdiff_right h.1 }⟩ + +@[simp] lemma sdiff_eq_left : x \ y = x ↔ disjoint x y := +⟨λ h, disjoint_sdiff_self_left.mono_left h.ge, disjoint.sdiff_eq_left⟩ + /- TODO: we could make an alternative constructor for `generalized_boolean_algebra` using `disjoint x (y \ x)` and `x ⊔ (y \ x) = y` as axioms. -/ theorem disjoint.sdiff_eq_of_sup_eq (hi : disjoint x z) (hs : x ⊔ z = y) : y \ x = z := diff --git a/src/order/rel_classes.lean b/src/order/rel_classes.lean index 4ebbf966321d7..b22650778f893 100644 --- a/src/order/rel_classes.lean +++ b/src/order/rel_classes.lean @@ -459,6 +459,8 @@ instance is_nonstrict_strict_order.to_is_irrefl {r : α → α → Prop} {s : α section subset variables [has_subset α] {a b c : α} +lemma subset_of_eq_of_subset (hab : a = b) (hbc : b ⊆ c) : a ⊆ c := by rwa hab +lemma subset_of_subset_of_eq (hab : a ⊆ b) (hbc : b = c) : a ⊆ c := by rwa ←hbc @[refl] lemma subset_refl [is_refl α (⊆)] (a : α) : a ⊆ a := refl _ lemma subset_rfl [is_refl α (⊆)] : a ⊆ a := refl _ lemma subset_of_eq [is_refl α (⊆)] : a = b → a ⊆ b := λ h, h ▸ subset_rfl @@ -473,6 +475,8 @@ antisymm h h' lemma superset_antisymm [is_antisymm α (⊆)] (h : a ⊆ b) (h' : b ⊆ a) : b = a := antisymm' h h' +alias subset_of_eq_of_subset ← eq.trans_subset +alias subset_of_subset_of_eq ← has_subset.subset.trans_eq alias subset_of_eq ← eq.subset' --TODO: Fix it and kill `eq.subset` alias superset_of_eq ← eq.superset alias subset_trans ← has_subset.subset.trans @@ -488,8 +492,10 @@ lemma superset_antisymm_iff [is_refl α (⊆)] [is_antisymm α (⊆)] : a = b end subset section ssubset -variables [has_ssubset α] +variables [has_ssubset α] {a b c : α} +lemma ssubset_of_eq_of_ssubset (hab : a = b) (hbc : b ⊂ c) : a ⊂ c := by rwa hab +lemma ssubset_of_ssubset_of_eq (hab : a ⊂ b) (hbc : b = c) : a ⊂ c := by rwa ←hbc lemma ssubset_irrefl [is_irrefl α (⊂)] (a : α) : ¬ a ⊂ a := irrefl _ lemma ssubset_irrfl [is_irrefl α (⊂)] {a : α} : ¬ a ⊂ a := irrefl _ lemma ne_of_ssubset [is_irrefl α (⊂)] {a b : α} : a ⊂ b → a ≠ b := ne_of_irrefl @@ -497,6 +503,8 @@ lemma ne_of_ssuperset [is_irrefl α (⊂)] {a b : α} : a ⊂ b → b ≠ a := n @[trans] lemma ssubset_trans [is_trans α (⊂)] {a b c : α} : a ⊂ b → b ⊂ c → a ⊂ c := trans lemma ssubset_asymm [is_asymm α (⊂)] {a b : α} (h : a ⊂ b) : ¬ b ⊂ a := asymm h +alias ssubset_of_eq_of_ssubset ← eq.trans_ssubset +alias ssubset_of_ssubset_of_eq ← has_ssubset.ssubset.trans_eq alias ssubset_irrfl ← has_ssubset.ssubset.false alias ne_of_ssubset ← has_ssubset.ssubset.ne alias ne_of_ssuperset ← has_ssubset.ssubset.ne' diff --git a/src/topology/instances/ennreal.lean b/src/topology/instances/ennreal.lean index ae2089995f426..dfe2b9fc47fd8 100644 --- a/src/topology/instances/ennreal.lean +++ b/src/topology/instances/ennreal.lean @@ -935,7 +935,7 @@ lemma tsum_union_le (f : α → ℝ≥0∞) (s t : set α) : calc ∑' (x : s ∪ t), f x = ∑' (x : s ∪ (t \ s)), f x : by { apply tsum_congr_subtype, rw union_diff_self } ... = ∑' (x : s), f x + ∑' (x : t \ s), f x : - tsum_union_disjoint disjoint_diff ennreal.summable ennreal.summable + tsum_union_disjoint disjoint_sdiff_self_right ennreal.summable ennreal.summable ... ≤ ∑' (x : s), f x + ∑' (x : t), f x : add_le_add le_rfl (tsum_mono_subtype _ (diff_subset _ _)) From 7c60702a4145fa8de521d53b93e602463f555cd7 Mon Sep 17 00:00:00 2001 From: loefflerd Date: Sun, 8 Jan 2023 09:48:25 +0000 Subject: [PATCH 0179/1291] feat(measure_theory/integral): "integral_prod_mul" for complex coefficients (#18085) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This generalises some lemmas in the integration library, currently only stated for `ℝ`-valued functions, to allow complex-valued functions too. (Note that one can't generalise all the way to an arbitrary Banach algebra, there are problems when one side is integrable and the other isn't.) --- src/measure_theory/constructions/prod.lean | 8 ++-- src/measure_theory/function/l1_space.lean | 52 +++++++++++----------- src/measure_theory/integral/bochner.lean | 11 +++-- src/probability/variance.lean | 4 +- 4 files changed, 40 insertions(+), 35 deletions(-) diff --git a/src/measure_theory/constructions/prod.lean b/src/measure_theory/constructions/prod.lean index d34a2ec880184..2f7ef21f9d5b7 100644 --- a/src/measure_theory/constructions/prod.lean +++ b/src/measure_theory/constructions/prod.lean @@ -904,7 +904,8 @@ lemma integrable.integral_norm_prod_right [sigma_finite μ] ⦃f : α × β → (hf : integrable f (μ.prod ν)) : integrable (λ y, ∫ x, ‖f (x, y)‖ ∂μ) ν := hf.swap.integral_norm_prod_left -lemma integrable_prod_mul {f : α → ℝ} {g : β → ℝ} (hf : integrable f μ) (hg : integrable g ν) : +lemma integrable_prod_mul {L : Type*} [is_R_or_C L] + {f : α → L} {g : β → L} (hf : integrable f μ) (hg : integrable g ν) : integrable (λ (z : α × β), f z.1 * g z.2) (μ.prod ν) := begin refine (integrable_prod_iff _).2 ⟨_, _⟩, @@ -1083,7 +1084,7 @@ begin exact integral_prod f hf end -lemma integral_prod_mul (f : α → ℝ) (g : β → ℝ) : +lemma integral_prod_mul {L : Type*} [is_R_or_C L] (f : α → L) (g : β → L) : ∫ z, f z.1 * g z.2 ∂(μ.prod ν) = (∫ x, f x ∂μ) * (∫ y, g y ∂ν) := begin by_cases h : integrable (λ (z : α × β), f z.1 * g z.2) (μ.prod ν), @@ -1096,7 +1097,8 @@ begin simp [integral_undef h, integral_undef H], end -lemma set_integral_prod_mul (f : α → ℝ) (g : β → ℝ) (s : set α) (t : set β) : +lemma set_integral_prod_mul {L : Type*} [is_R_or_C L] + (f : α → L) (g : β → L) (s : set α) (t : set β) : ∫ z in s ×ˢ t, f z.1 * g z.2 ∂(μ.prod ν) = (∫ x in s, f x ∂μ) * (∫ y in t, g y ∂ν) := by simp only [← measure.prod_restrict s t, integrable_on, integral_prod_mul] diff --git a/src/measure_theory/function/l1_space.lean b/src/measure_theory/function/l1_space.lean index 80b6a243cb362..9caa52b58626c 100644 --- a/src/measure_theory/function/l1_space.lean +++ b/src/measure_theory/function/l1_space.lean @@ -882,27 +882,47 @@ lemma integrable_smul_iff {c : 𝕜} (hc : c ≠ 0) (f : α → β) : integrable (c • f) μ ↔ integrable f μ := and_congr (ae_strongly_measurable_const_smul_iff₀ hc) (has_finite_integral_smul_iff hc f) -lemma integrable.const_mul {f : α → ℝ} (h : integrable f μ) (c : ℝ) : +end normed_space + +section normed_space_over_complete_field +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] [complete_space 𝕜] +variables {E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] + +lemma integrable_smul_const {f : α → 𝕜} {c : E} (hc : c ≠ 0) : + integrable (λ x, f x • c) μ ↔ integrable f μ := +begin + simp_rw [integrable, ae_strongly_measurable_smul_const_iff hc, and.congr_right_iff, + has_finite_integral, nnnorm_smul, ennreal.coe_mul], + intro hf, rw [lintegral_mul_const' _ _ ennreal.coe_ne_top, ennreal.mul_lt_top_iff], + have : ∀ x : ℝ≥0∞, x = 0 → x < ∞ := by simp, + simp [hc, or_iff_left_of_imp (this _)] +end +end normed_space_over_complete_field + +section is_R_or_C +variables {𝕜 : Type*} [is_R_or_C 𝕜] {f : α → 𝕜} + +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 : ℝ) : +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 : ℝ) : +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 : ℝ) : +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 : ℝ) : +lemma integrable.div_const {f : α → 𝕜} (h : integrable f μ) (c : 𝕜) : integrable (λ x, f x / c) μ := by simp_rw [div_eq_mul_inv, h.mul_const] -lemma integrable.bdd_mul' {f g : α → ℝ} {c : ℝ} (hg : integrable g μ) +lemma integrable.bdd_mul' {f g : α → 𝕜} {c : ℝ} (hg : integrable g μ) (hf : ae_strongly_measurable f μ) (hf_bound : ∀ᵐ x ∂μ, ‖f x‖ ≤ c) : integrable (λ x, f x * g x) μ := begin @@ -912,26 +932,6 @@ begin exact (norm_mul_le _ _).trans (mul_le_mul_of_nonneg_right hx (norm_nonneg _)), end -end normed_space - -section normed_space_over_complete_field -variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] [complete_space 𝕜] -variables {E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] - -lemma integrable_smul_const {f : α → 𝕜} {c : E} (hc : c ≠ 0) : - integrable (λ x, f x • c) μ ↔ integrable f μ := -begin - simp_rw [integrable, ae_strongly_measurable_smul_const_iff hc, and.congr_right_iff, - has_finite_integral, nnnorm_smul, ennreal.coe_mul], - intro hf, rw [lintegral_mul_const' _ _ ennreal.coe_ne_top, ennreal.mul_lt_top_iff], - have : ∀ x : ℝ≥0∞, x = 0 → x < ∞ := by simp, - simp [hc, or_iff_left_of_imp (this _)] -end -end normed_space_over_complete_field - -section is_R_or_C -variables {𝕜 : Type*} [is_R_or_C 𝕜] {f : α → 𝕜} - lemma integrable.of_real {f : α → ℝ} (hf : integrable f μ) : integrable (λ x, (f x : 𝕜)) μ := by { rw ← mem_ℒp_one_iff_integrable at hf ⊢, exact hf.of_real } diff --git a/src/measure_theory/integral/bochner.lean b/src/measure_theory/integral/bochner.lean index 5fcec00ea6050..a21333f01c18a 100644 --- a/src/measure_theory/integral/bochner.lean +++ b/src/measure_theory/integral/bochner.lean @@ -770,14 +770,17 @@ lemma integral_smul (c : 𝕜) (f : α → E) : ∫ a, c • (f a) ∂μ = c • ∫ a, f a ∂μ := set_to_fun_smul (dominated_fin_meas_additive_weighted_smul μ) weighted_smul_smul c f -lemma integral_mul_left (r : ℝ) (f : α → ℝ) : ∫ a, r * (f a) ∂μ = r * ∫ a, f a ∂μ := +lemma integral_mul_left {L : Type*} [is_R_or_C L] (r : L) (f : α → L) : + ∫ a, r * (f a) ∂μ = r * ∫ a, f a ∂μ := integral_smul r f -lemma integral_mul_right (r : ℝ) (f : α → ℝ) : ∫ a, (f a) * r ∂μ = ∫ a, f a ∂μ * r := +lemma integral_mul_right {L : Type*} [is_R_or_C L] (r : L) (f : α → L) : + ∫ a, (f a) * r ∂μ = ∫ a, f a ∂μ * r := by { simp only [mul_comm], exact integral_mul_left r f } -lemma integral_div (r : ℝ) (f : α → ℝ) : ∫ a, (f a) / r ∂μ = ∫ a, f a ∂μ / r := -integral_mul_right r⁻¹ f +lemma integral_div {L : Type*} [is_R_or_C L] (r : L) (f : α → L) : + ∫ a, (f a) / r ∂μ = ∫ a, f a ∂μ / r := +by simpa only [←div_eq_mul_inv] using integral_mul_right r⁻¹ f lemma integral_congr_ae (h : f =ᵐ[μ] g) : ∫ a, f a ∂μ = ∫ a, g a ∂μ := set_to_fun_congr_ae (dominated_fin_meas_additive_weighted_smul μ) h diff --git a/src/probability/variance.lean b/src/probability/variance.lean index 8b67daebd8f56..2e539b0e9810e 100644 --- a/src/probability/variance.lean +++ b/src/probability/variance.lean @@ -362,7 +362,7 @@ begin { apply mem_ℒp.integrable_sq, exact mem_ℒp_finset_sum' _ (λ i hi, (hs _ (mem_insert_of_mem hi))) } }, { rw mul_assoc, - apply integrable.const_mul _ 2, + apply integrable.const_mul _ (2:ℝ), simp only [mul_sum, sum_apply, pi.mul_apply], apply integrable_finset_sum _ (λ i hi, _), apply indep_fun.integrable_mul _ @@ -383,7 +383,7 @@ begin simp only [mul_assoc, integral_mul_left, pi.mul_apply, pi.bit0_apply, pi.one_apply, sum_apply, add_right_eq_self, mul_sum], rw integral_finset_sum s (λ i hi, _), swap, - { apply integrable.const_mul _ 2, + { apply integrable.const_mul _ (2:ℝ), apply indep_fun.integrable_mul _ (mem_ℒp.integrable one_le_two (hs _ (mem_insert_self _ _))) (mem_ℒp.integrable one_le_two (hs _ (mem_insert_of_mem hi))), From 247a102b14f3cebfee126293341af5f6bed00237 Mon Sep 17 00:00:00 2001 From: loefflerd Date: Sun, 8 Jan 2023 13:02:49 +0000 Subject: [PATCH 0180/1291] feat(analysis/fourier): convergence of Fourier series (#17913) This PR adds a straightforward but useful criterion for convergence of Fourier series: for a continuous periodic function `f`, if the sequence of Fourier coefficients of `f` is absolutely summable, then the Fourier series converges uniformly, and hence pointwise everywhere, to `f`. (Note that it is obvious in this case that the Fourier series converges uniformly to _something_, the difficult part is that the limit is actually `f`.) Co-authored-by: Junyan Xu --- docs/100.yaml | 4 +- docs/undergrad.yaml | 6 +- src/analysis/fourier/add_circle.lean | 180 ++++++++++++------ src/measure_theory/function/l2_space.lean | 4 +- src/measure_theory/function/lp_space.lean | 50 ++++- src/measure_theory/integral/periodic.lean | 75 ++++++-- src/topology/continuous_function/compact.lean | 7 + 7 files changed, 236 insertions(+), 90 deletions(-) diff --git a/docs/100.yaml b/docs/100.yaml index dde7c7548708f..a14251360748f 100644 --- a/docs/100.yaml +++ b/docs/100.yaml @@ -297,8 +297,8 @@ 76: title : Fourier Series decls : - - fourier_series_repr - - has_sum_fourier_series + - fourier_coeff + - has_sum_fourier_series_L2 author : Heather Macbeth 77: title : Sum of kth powers diff --git a/docs/undergrad.yaml b/docs/undergrad.yaml index b775dd5c45787..18d92c12e930f 100644 --- a/docs/undergrad.yaml +++ b/docs/undergrad.yaml @@ -436,7 +436,7 @@ Topology: inner product space $L^2$: 'measure_theory.L2.inner_product_space' its completeness: 'measure_theory.Lp.complete_space' Hilbert bases: 'hilbert_basis' # the document specifies "in the separable case" but we don't require this - example, the Hilbert basis of trigonometric polynomials: 'fourier_series' + example, the Hilbert basis of trigonometric polynomials: 'fourier_basis' example, classical Hilbert bases of orthogonal polynomials: '' Lax-Milgram theorem: 'is_coercive.continuous_linear_equiv_of_bilin' $H^1_0([0,1])$ and its application to the one-dimensional Dirichlet problem: '' @@ -521,12 +521,12 @@ Measures and integral calculus: approximation by convolution: 'cont_diff_bump_of_inner.convolution_tendsto_right' regularization by convolution: 'has_compact_support.cont_diff_convolution_left' Fourier analysis: - Fourier series of locally integrable periodic real-valued functions: 'fourier_series' + Fourier series of locally integrable periodic real-valued functions: 'fourier_basis' Riemann-Lebesgue lemma: '' convolution product of periodic functions: '' Dirichlet theorem: '' Fejer theorem: '' - Parseval theorem: 'tsum_sq_fourier_series_repr' + Parseval theorem: 'tsum_sq_fourier_coeff' Fourier transforms on $\mathrm{L}^1(\R^d)$ and $\mathrm{L}^2(\R^d)$: '' Plancherel’s theorem: '' diff --git a/src/analysis/fourier/add_circle.lean b/src/analysis/fourier/add_circle.lean index b4b0ece6ab1a5..0a24bd245014b 100644 --- a/src/analysis/fourier/add_circle.lean +++ b/src/analysis/fourier/add_circle.lean @@ -26,36 +26,34 @@ This file contains basic results on Fourier series for functions on the additive so we do not declare it as a `measure_space` instance, to avoid confusion.) * for `n : ℤ`, `fourier n` is the monomial `λ x, exp (2 π i n x / T)`, bundled as a continuous map from `add_circle T` to `ℂ`. -* for `n : ℤ` and `p : ℝ≥0∞`, `fourier_Lp p n` is an abbreviation for the monomial `fourier n` - considered as an element of the Lᵖ-space `Lp ℂ p haar_add_circle`, via the embedding - `continuous_map.to_Lp`. -* `fourier_series` is the canonical isometric isomorphism from `Lp ℂ 2 haar_add_circle` to - `ℓ²(ℤ, ℂ)` induced by taking Fourier coefficients. +* `fourier_basis` is the Hilbert basis of `Lp ℂ 2 haar_add_circle` given by the images of the + monomials `fourier n`. +* `fourier_coeff f n`, for `f : add_circle T → ℂ`, is the `n`-th Fourier coefficient of `f` + (defined as an integral over `add_circle T`). ## Main statements The theorem `span_fourier_closure_eq_top` states that the span of the monomials `fourier n` is dense in `C(add_circle T, ℂ)`, i.e. that its `submodule.topological_closure` is `⊤`. This follows -from the Stone-Weierstrass theorem after checking that it is a subalgebra, closed under conjugation, -and separates points. - -The theorem `span_fourier_Lp_closure_eq_top` states that for `1 ≤ p < ∞` the span of the monomials -`fourier_Lp` is dense in the Lᵖ space of `add_circle T`, i.e. that its -`submodule.topological_closure` is `⊤`. This follows from the previous theorem using general theory -on approximation of Lᵖ functions by continuous functions. - -The theorem `orthonormal_fourier` states that the monomials `fourier_Lp 2 n` form an orthonormal set -(in the L² space of `add_circle T` with respect to `haar_add_circle`). - -The last two results together provide that the functions `fourier_Lp 2 n` form a Hilbert basis for -L²; this is named as `fourier_series`. - -Parseval's identity, `tsum_sq_fourier_series_repr`, is a direct consequence of the construction of -this Hilbert basis. +from the Stone-Weierstrass theorem after checking that the span is a subalgebra, is closed under +conjugation, and separates points. + +Using this and general theory on approximation of Lᵖ functions by continuous functions, we deduce +(`span_fourier_Lp_closure_eq_top`) that for any `1 ≤ p < ∞`, the span of the Fourier monomials is +dense in the Lᵖ space of `add_circle T`. For `p = 2` we show (`orthonormal_fourier`) that the +monomials are also orthonormal, so they form a Hilbert basis for L², which is named as +`fourier_basis`; in particular, for `L²` functions `f`, the Fourier series of `f` converges to `f` +in the `L²` topology (`has_sum_fourier_series_L2`). Parseval's identity, `tsum_sq_fourier_coeff`, is +a direct consequence. + +For continuous maps `f : add_circle T → ℂ`, the theorem +`continuous_map.has_sum_fourier_series_of_summable` states that if the sequence of Fourier +coefficients of `f` is summable, then the Fourier series `∑ (i:ℤ), f.fourier_coeff i * fourier i` +converges to `f` in the uniform-convergence topology of `C(add_circle T, ℂ)`. -/ noncomputable theory -open_locale ennreal complex_conjugate classical real +open_locale ennreal complex_conjugate real open topological_space continuous_map measure_theory measure_theory.measure algebra submodule set variables {T : ℝ} @@ -135,13 +133,27 @@ def fourier (n : ℤ) : C(add_circle T, ℂ) := @[simp] lemma fourier_apply {n : ℤ} {x : add_circle T} : fourier n x = to_circle (n • x) := rfl +@[simp] lemma fourier_coe_apply {n : ℤ} {x : ℝ} : + fourier n (x : add_circle T) = complex.exp (2 * π * complex.I * n * x / T) := +begin + rw [fourier_apply, ←quotient_add_group.coe_zsmul, to_circle, function.periodic.lift_coe, + exp_map_circle_apply, complex.of_real_mul, complex.of_real_div, complex.of_real_mul, + zsmul_eq_mul, complex.of_real_mul, complex.of_real_int_cast, complex.of_real_bit0, + complex.of_real_one], + congr' 1, ring, +end + @[simp] lemma fourier_zero {x : add_circle T} : fourier 0 x = 1 := begin induction x using quotient_add_group.induction_on', - rw [fourier_apply, to_circle, zero_zsmul, ←quotient_add_group.coe_zero, - function.periodic.lift_coe, mul_zero, exp_map_circle_zero, coe_one_unit_sphere], + simp only [fourier_coe_apply, algebra_map.coe_zero, mul_zero, zero_mul, + zero_div, complex.exp_zero], end +@[simp] lemma fourier_eval_zero (n : ℤ) : fourier n (0 : add_circle T) = 1 := +by rw [←quotient_add_group.coe_zero, fourier_coe_apply, complex.of_real_zero, + mul_zero, zero_div, complex.exp_zero] + @[simp] lemma fourier_one {x : add_circle T} : fourier 1 x = to_circle x := by rw [fourier_apply, one_zsmul] @@ -153,9 +165,17 @@ begin end @[simp] lemma fourier_add {m n : ℤ} {x : add_circle T} : - fourier (m + n) x = (fourier m x) * (fourier n x) := + fourier (m + n) x = fourier m x * fourier n x := by simp_rw [fourier_apply, add_zsmul, to_circle_add, coe_mul_unit_sphere] +lemma fourier_norm [fact (0 < T)] (n : ℤ) : ‖@fourier T n‖ = 1 := +begin + rw continuous_map.norm_eq_supr_norm, + have : ∀ (x : add_circle T), ‖fourier n x‖ = 1 := λ x, abs_coe_circle _, + simp_rw this, + exact @csupr_const _ _ _ has_zero.nonempty _, +end + /-- For `n ≠ 0`, a translation by `T / 2 / n` negates the function `fourier n`. -/ lemma fourier_add_half_inv_index {n : ℤ} (hn : n ≠ 0) (hT : 0 < T) (x : add_circle T) : fourier n (x + ((T / 2 / n) : ℝ)) = - fourier n x := @@ -266,46 +286,74 @@ end end monomials -section fourier +section scope_hT -- everything from here on needs `0 < T` variables [hT : fact (0 < T)] include hT -/-- We define `fourier_series` to be a `ℤ`-indexed Hilbert basis for `Lp ℂ 2 haar_add_circle`, + +section fourier_coeff +variables {E : Type} [normed_add_comm_group E] [normed_space ℂ E] [complete_space E] + +/-- The `n`-th Fourier coefficient of a function `add_circle T → E`, for `E` a complete normed +`ℂ`-vector space, defined as the integral over `add_circle T` of `fourier (-n) t • f t`. -/ +def fourier_coeff (f : add_circle T → E) (n : ℤ) : E := +∫ (t : add_circle T), fourier (-n) t • f t ∂ haar_add_circle + +/-- The Fourier coefficients of a function can be computed as an integral +over `[a, a + T]` for any real `a`. -/ +lemma fourier_coeff_eq_interval_integral (f : add_circle T → E) (n : ℤ) (a : ℝ) : + fourier_coeff f n = (1 / T) • ∫ x in a .. a + T, @fourier T (-n) x • f x := +begin + have : ∀ (x : ℝ), @fourier T (-n) x • f x = (λ (z : add_circle T), @fourier T (-n) z • f z) x, + { intro x, refl, }, + simp_rw this, + rw [fourier_coeff, add_circle.interval_integral_preimage T a, + volume_eq_smul_haar_add_circle, integral_smul_measure, ennreal.to_real_of_real hT.out.le, + ←smul_assoc, smul_eq_mul, one_div_mul_cancel hT.out.ne', one_smul], +end + +end fourier_coeff + +section fourier_L2 + +/-- We define `fourier_basis` to be a `ℤ`-indexed Hilbert basis for `Lp ℂ 2 haar_add_circle`, which by definition is an isometric isomorphism from `Lp ℂ 2 haar_add_circle` to `ℓ²(ℤ, ℂ)`. -/ -def fourier_series : hilbert_basis ℤ ℂ (Lp ℂ 2 $ @haar_add_circle T hT) := +def fourier_basis : hilbert_basis ℤ ℂ (Lp ℂ 2 $ @haar_add_circle T hT) := hilbert_basis.mk orthonormal_fourier (span_fourier_Lp_closure_eq_top (by norm_num)).ge -/-- The elements of the Hilbert basis `fourier_series` are the functions `fourier_Lp 2`, i.e. the +/-- The elements of the Hilbert basis `fourier_basis` are the functions `fourier_Lp 2`, i.e. the monomials `fourier n` on the circle considered as elements of `L²`. -/ -@[simp] lemma coe_fourier_series : ⇑(@fourier_series _ hT) = fourier_Lp 2:= hilbert_basis.coe_mk _ _ +@[simp] lemma coe_fourier_basis : ⇑(@fourier_basis _ hT) = fourier_Lp 2 := hilbert_basis.coe_mk _ _ -/-- Under the isometric isomorphism `fourier_series` from `Lp ℂ 2 haar_circle` to `ℓ²(ℤ, ℂ)`, the -`i`-th coefficient is the integral over `add_circle T` of `λ t, fourier (-i) t * f t`. -/ -lemma fourier_series_repr (f : Lp ℂ 2 $ @haar_add_circle T hT) (i : ℤ) : - fourier_series.repr f i = ∫ (t : add_circle T), fourier (-i) t * f t ∂ haar_add_circle := +/-- Under the isometric isomorphism `fourier_basis` from `Lp ℂ 2 haar_circle` to `ℓ²(ℤ, ℂ)`, the +`i`-th coefficient is `fourier_coeff f i`, i.e., the integral over `add_circle T` of +`λ t, fourier (-i) t * f t` with respect to the Haar measure of total mass 1. -/ +lemma fourier_basis_repr (f : Lp ℂ 2 $ @haar_add_circle T hT) (i : ℤ) : + fourier_basis.repr f i = fourier_coeff f i := begin transitivity ∫ (t : add_circle T), conj (((@fourier_Lp T hT 2 _ i) : add_circle T → ℂ) t) * f t ∂ haar_add_circle, - { simp [fourier_series.repr_apply_apply f i, measure_theory.L2.inner_def] }, + { simp [fourier_basis.repr_apply_apply f i, measure_theory.L2.inner_def] }, { apply integral_congr_ae, filter_upwards [coe_fn_fourier_Lp 2 i] with _ ht, - rw [ht, ←fourier_neg] } + rw [ht, ←fourier_neg, smul_eq_mul], } end /-- The Fourier series of an `L2` function `f` sums to `f`, in the `L²` space of `add_circle T`. -/ -lemma has_sum_fourier_series (f : Lp ℂ 2 $ @haar_add_circle T hT) : - has_sum (λ i, fourier_series.repr f i • fourier_Lp 2 i) f := -by simpa using hilbert_basis.has_sum_repr fourier_series f +lemma has_sum_fourier_series_L2 (f : Lp ℂ 2 $ @haar_add_circle T hT) : + has_sum (λ i, fourier_coeff f i • fourier_Lp 2 i) f := +by { simp_rw ←fourier_basis_repr, simpa using hilbert_basis.has_sum_repr fourier_basis f } /-- **Parseval's identity**: for an `L²` function `f` on `add_circle T`, the sum of the squared norms of the Fourier coefficients equals the `L²` norm of `f`. -/ -lemma tsum_sq_fourier_series_repr (f : Lp ℂ 2 $ @haar_add_circle T hT) : - ∑' i : ℤ, ‖fourier_series.repr f i‖ ^ 2 = ∫ (t : add_circle T), ‖f t‖ ^ 2 ∂ haar_add_circle := +lemma tsum_sq_fourier_coeff (f : Lp ℂ 2 $ @haar_add_circle T hT) : + ∑' i : ℤ, ‖fourier_coeff f i‖ ^ 2 = ∫ (t : add_circle T), ‖f t‖ ^ 2 ∂ haar_add_circle := begin - have H₁ : ‖fourier_series.repr f‖ ^ 2 = ∑' i, ‖fourier_series.repr f i‖ ^ 2, - { exact_mod_cast lp.norm_rpow_eq_tsum _ (fourier_series.repr f), + simp_rw ←fourier_basis_repr, + have H₁ : ‖fourier_basis.repr f‖ ^ 2 = ∑' i, ‖fourier_basis.repr f i‖ ^ 2, + { exact_mod_cast lp.norm_rpow_eq_tsum _ (fourier_basis.repr f), norm_num }, - have H₂ : ‖fourier_series.repr f‖ ^ 2 = ‖f‖ ^ 2 := by simp, + have H₂ : ‖fourier_basis.repr f‖ ^ 2 = ‖f‖ ^ 2 := by simp, have H₃ := congr_arg is_R_or_C.re (@L2.inner_def (add_circle T) ℂ ℂ _ _ _ _ f f), rw ← integral_re at H₃, { simp only [← norm_sq_eq_inner] at H₃, @@ -313,17 +361,39 @@ begin { exact L2.integrable_inner f f }, end -/-- The Fourier coefficients are given by integrating over the interval `[a, a + T] ⊂ ℝ`. -/ -lemma fourier_series_repr' (f : Lp ℂ 2 $ @haar_add_circle T hT) (n : ℤ) (a : ℝ): - fourier_series.repr f n = 1 / T * ∫ x in a .. a + T, @fourier T (-n) x * f x := +end fourier_L2 + +section convergence + +variables (f : C(add_circle T, ℂ)) + +lemma fourier_coeff_to_Lp (n : ℤ) : + fourier_coeff (to_Lp 2 haar_add_circle ℂ f) n = fourier_coeff f n := +integral_congr_ae (filter.eventually_eq.mul + (filter.eventually_of_forall (by tauto)) + (continuous_map.coe_fn_to_ae_eq_fun haar_add_circle f)) + +variables {f} + +/-- If the sequence of Fourier coefficients of `f` is summable, then the Fourier series converges +uniformly to `f`. -/ +lemma has_sum_fourier_series_of_summable (h : summable (fourier_coeff f)) : + has_sum (λ i, fourier_coeff f i • fourier i) f := begin - have ha : ae_strongly_measurable (λ (t : add_circle T), fourier (-n) t * f t) haar_add_circle := - (map_continuous _).ae_strongly_measurable.mul (Lp.ae_strongly_measurable _), - rw [fourier_series_repr, add_circle.interval_integral_preimage T a (ha.smul_measure _), - volume_eq_smul_haar_add_circle, integral_smul_measure], - have : (T : ℂ) ≠ 0 := by exact_mod_cast hT.out.ne', - field_simp [ennreal.to_real_of_real hT.out.le, complex.real_smul], - ring, + have sum_L2 := has_sum_fourier_series_L2 (to_Lp 2 haar_add_circle ℂ f), + simp_rw fourier_coeff_to_Lp at sum_L2, + refine continuous_map.has_sum_of_has_sum_Lp (summable_of_summable_norm _) sum_L2, + simp_rw [norm_smul, fourier_norm, mul_one, summable_norm_iff], + exact h, end -end fourier +/-- If the sequence of Fourier coefficients of `f` is summable, then the Fourier series of `f` +converges everywhere pointwise to `f`. -/ +lemma has_pointwise_sum_fourier_series_of_summable + (h : summable (fourier_coeff f)) (x : add_circle T) : + has_sum (λ i, fourier_coeff f i • fourier i x) (f x) := +(continuous_map.eval_clm ℂ x).has_sum (has_sum_fourier_series_of_summable h) + +end convergence + +end scope_hT diff --git a/src/measure_theory/function/l2_space.lean b/src/measure_theory/function/l2_space.lean index f865a5f5c37e5..04fa719cde5b8 100644 --- a/src/measure_theory/function/l2_space.lean +++ b/src/measure_theory/function/l2_space.lean @@ -236,8 +236,8 @@ lemma bounded_continuous_function.inner_to_Lp (f g : α →ᵇ 𝕜) : = ∫ x, conj (f x) * g x ∂μ := begin apply integral_congr_ae, - have hf_ae := f.coe_fn_to_Lp μ, - have hg_ae := g.coe_fn_to_Lp μ, + have hf_ae := f.coe_fn_to_Lp 2 μ 𝕜, + have hg_ae := g.coe_fn_to_Lp 2 μ 𝕜, filter_upwards [hf_ae, hg_ae] with _ hf hg, rw [hf, hg], simp diff --git a/src/measure_theory/function/lp_space.lean b/src/measure_theory/function/lp_space.lean index 143ea4968c7be..9c98bf1ea7d94 100644 --- a/src/measure_theory/function/lp_space.lean +++ b/src/measure_theory/function/lp_space.lean @@ -2826,11 +2826,11 @@ begin (by { rintros - ⟨f, rfl⟩, exact mem_Lp f } : _ ≤ Lp E p μ), end -variables (𝕜 : Type*) +variables (𝕜 : Type*) [fact (1 ≤ p)] /-- The bounded linear map of considering a bounded continuous function on a finite-measure space as an element of `Lp`. -/ -def to_Lp [normed_field 𝕜] [normed_space 𝕜 E] [fact (1 ≤ p)] : +def to_Lp [normed_field 𝕜] [normed_space 𝕜 E] : (α →ᵇ E) →L[𝕜] (Lp E p μ) := linear_map.mk_continuous (linear_map.cod_restrict @@ -2840,23 +2840,34 @@ linear_map.mk_continuous _ Lp_norm_le +lemma coe_fn_to_Lp [normed_field 𝕜] [normed_space 𝕜 E] (f : α →ᵇ E) : + to_Lp p μ 𝕜 f =ᵐ[μ] f := ae_eq_fun.coe_fn_mk f _ + variables {𝕜} -lemma range_to_Lp [normed_field 𝕜] [normed_space 𝕜 E] [fact (1 ≤ p)] : +lemma range_to_Lp [normed_field 𝕜] [normed_space 𝕜 E] : ((linear_map.range (to_Lp p μ 𝕜 : (α →ᵇ E) →L[𝕜] Lp E p μ)).to_add_subgroup) = measure_theory.Lp.bounded_continuous_function E p μ := range_to_Lp_hom p μ variables {p} -lemma coe_fn_to_Lp [normed_field 𝕜] [normed_space 𝕜 E] [fact (1 ≤ p)] (f : α →ᵇ E) : - to_Lp p μ 𝕜 f =ᵐ[μ] f := -ae_eq_fun.coe_fn_mk f _ - -lemma to_Lp_norm_le [nontrivially_normed_field 𝕜] [normed_space 𝕜 E] [fact (1 ≤ p)] : +lemma to_Lp_norm_le [nontrivially_normed_field 𝕜] [normed_space 𝕜 E]: ‖(to_Lp p μ 𝕜 : (α →ᵇ E) →L[𝕜] (Lp E p μ))‖ ≤ (measure_univ_nnreal μ) ^ (p.to_real)⁻¹ := linear_map.mk_continuous_norm_le _ ((measure_univ_nnreal μ) ^ (p.to_real)⁻¹).coe_nonneg _ +lemma to_Lp_inj {f g : α →ᵇ E} [μ.is_open_pos_measure] [normed_field 𝕜] [normed_space 𝕜 E] : + to_Lp p μ 𝕜 f = to_Lp p μ 𝕜 g ↔ f = g := +begin + refine ⟨λ h, _, by tauto⟩, + rw [←fun_like.coe_fn_eq, ←(map_continuous f).ae_eq_iff_eq μ (map_continuous g)], + refine (coe_fn_to_Lp p μ 𝕜 f).symm.trans (eventually_eq.trans _ $ coe_fn_to_Lp p μ 𝕜 g), + rw h, +end + +lemma to_Lp_injective [μ.is_open_pos_measure] [normed_field 𝕜] [normed_space 𝕜 E] : + function.injective ⇑(to_Lp p μ 𝕜 : (α →ᵇ E) →L[𝕜] (Lp E p μ)) := λ f g hfg, (to_Lp_inj μ).mp hfg + end bounded_continuous_function namespace continuous_map @@ -2906,7 +2917,28 @@ rfl (to_Lp p μ 𝕜 f : α →ₘ[μ] E) = f.to_ae_eq_fun μ := rfl -variables [nontrivially_normed_field 𝕜] [normed_space 𝕜 E] +lemma to_Lp_injective [μ.is_open_pos_measure] [normed_field 𝕜] [normed_space 𝕜 E] : + function.injective ⇑(to_Lp p μ 𝕜 : C(α, E) →L[𝕜] (Lp E p μ)) := +(bounded_continuous_function.to_Lp_injective _).comp + (linear_isometry_bounded_of_compact α E 𝕜).injective + +lemma to_Lp_inj {f g : C(α, E)} [μ.is_open_pos_measure] [normed_field 𝕜] [normed_space 𝕜 E] : + to_Lp p μ 𝕜 f = to_Lp p μ 𝕜 g ↔ f = g := +(to_Lp_injective μ).eq_iff + +variables {μ} + +/-- If a sum of continuous functions `g n` is convergent, and the same sum converges in `Lᵖ` to `h`, +then in fact `g n` converges uniformly to `h`. -/ +lemma has_sum_of_has_sum_Lp {β : Type*} [μ.is_open_pos_measure] [normed_field 𝕜] [normed_space 𝕜 E] + {g : β → C(α, E)} {f : C(α, E)} (hg : summable g) + (hg2 : has_sum (to_Lp p μ 𝕜 ∘ g) (to_Lp p μ 𝕜 f)) : has_sum g f := +begin + convert summable.has_sum hg, + exact to_Lp_injective μ (hg2.unique ((to_Lp p μ 𝕜).has_sum $ summable.has_sum hg)), +end + +variables (μ) [nontrivially_normed_field 𝕜] [normed_space 𝕜 E] lemma to_Lp_norm_eq_to_Lp_norm_coe : ‖(to_Lp p μ 𝕜 : C(α, E) →L[𝕜] (Lp E p μ))‖ diff --git a/src/measure_theory/integral/periodic.lean b/src/measure_theory/integral/periodic.lean index dda056f622986..9f84cf6475e36 100644 --- a/src/measure_theory/integral/periodic.lean +++ b/src/measure_theory/integral/periodic.lean @@ -117,32 +117,71 @@ begin exact min_le_min (by linarith [hT.out]) (le_refl _), end -/-- The integral of a measurable function over `add_circle T` is equal to the integral over an +/-- The isomorphism `add_circle T ≃ Ioc a (a + T)` whose inverse is the natural quotient map, + as an equivalence of measurable spaces. -/ +noncomputable def measurable_equiv_Ioc (a : ℝ) : add_circle T ≃ᵐ Ioc a (a + T) := +{ measurable_to_fun := measurable_of_measurable_on_compl_singleton _ + (continuous_on_iff_continuous_restrict.mp $ continuous_at.continuous_on $ + λ x hx, continuous_at_equiv_Ioc T a hx).measurable, + measurable_inv_fun := add_circle.measurable_mk'.comp measurable_subtype_coe, + .. equiv_Ioc T a } + +/-- The isomorphism `add_circle T ≃ Ico a (a + T)` whose inverse is the natural quotient map, + as an equivalence of measurable spaces. -/ +noncomputable def measurable_equiv_Ico (a : ℝ) : add_circle T ≃ᵐ Ico a (a + T) := +{ measurable_to_fun := measurable_of_measurable_on_compl_singleton _ + (continuous_on_iff_continuous_restrict.mp $ continuous_at.continuous_on $ + λ x hx, continuous_at_equiv_Ico T a hx).measurable, + measurable_inv_fun := add_circle.measurable_mk'.comp measurable_subtype_coe, + .. equiv_Ico T a } + +/-- The lower integral of a function over `add_circle T` is equal to the lower integral over an interval (t, t + T] in `ℝ` of its lift to `ℝ`. -/ -protected lemma lintegral_preimage (t : ℝ) {f : add_circle T → ℝ≥0∞} (hf : measurable f) : +protected lemma lintegral_preimage (t : ℝ) (f : add_circle T → ℝ≥0∞) : ∫⁻ a in Ioc t (t + T), f a = ∫⁻ b : add_circle T, f b := -by rw [← lintegral_map hf add_circle.measurable_mk', - (add_circle.measure_preserving_mk T t).map_eq] +begin + have m : measurable_set (Ioc t (t + T)) := measurable_set_Ioc, + have := lintegral_map_equiv f (measurable_equiv_Ioc T t).symm, + swap, exact volume, + simp only [measurable_equiv_Ioc, equiv_Ioc, quotient_add_group.equiv_Ioc_mod, + measurable_equiv.symm_mk, measurable_equiv.coe_mk, equiv.coe_fn_symm_mk] at this, + rw ←(add_circle.measure_preserving_mk T t).map_eq, + convert this.symm using 1, -- TODO : there is no "set_lintegral_eq_subtype"? + { rw ←(map_comap_subtype_coe m _), + exact measurable_embedding.lintegral_map (measurable_embedding.subtype_coe m) _, }, + { congr' 1, + have : (coe : Ioc t (t + T) → add_circle T) = (coe : ℝ → add_circle T) ∘ (coe : _ → ℝ), + { ext1 x, refl, }, + simp_rw [this, ←map_map add_circle.measurable_mk' measurable_subtype_coe, + ←map_comap_subtype_coe m], + refl, } +end variables {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] [complete_space E] /-- The integral of an almost-everywhere strongly measurable function over `add_circle T` is equal to the integral over an interval (t, t + T] in `ℝ` of its lift to `ℝ`. -/ -protected lemma integral_preimage (t : ℝ) {f : add_circle T → E} - (hf : ae_strongly_measurable f volume) : +protected lemma integral_preimage (t : ℝ) (f : add_circle T → E) : ∫ a in Ioc t (t + T), f a = ∫ b : add_circle T, f b := begin - rw ← (add_circle.measure_preserving_mk T t).map_eq at ⊢ hf, - rw integral_map add_circle.measurable_mk'.ae_measurable hf, + have m : measurable_set (Ioc t (t + T)) := measurable_set_Ioc, + have := integral_map_equiv (measurable_equiv_Ioc T t).symm f, + simp only [measurable_equiv_Ioc, equiv_Ioc, quotient_add_group.equiv_Ioc_mod, + measurable_equiv.symm_mk, measurable_equiv.coe_mk, equiv.coe_fn_symm_mk, coe_coe] at this, + rw [←(add_circle.measure_preserving_mk T t).map_eq, set_integral_eq_subtype m, ←this], + have : (coe : Ioc t (t + T) → add_circle T) = (coe : ℝ → add_circle T) ∘ (coe : _ → ℝ), + { ext1 x, refl, }, + simp_rw [this, ←map_map add_circle.measurable_mk' measurable_subtype_coe, + ←map_comap_subtype_coe m], + refl, end /-- The integral of an almost-everywhere strongly measurable function over `add_circle T` is equal to the integral over an interval (t, t + T] in `ℝ` of its lift to `ℝ`. -/ -protected lemma interval_integral_preimage (t : ℝ) {f : add_circle T → E} - (hf : ae_strongly_measurable f volume) : +protected lemma interval_integral_preimage (t : ℝ) (f : add_circle T → E) : ∫ a in t..(t + T), f a = ∫ b : add_circle T, f b := begin - rw [integral_of_le, add_circle.integral_preimage T t hf], + rw [integral_of_le, add_circle.integral_preimage T t f], linarith [hT.out], end @@ -169,25 +208,23 @@ add_circle.measure_preserving_mk 1 t /-- The integral of a measurable function over `unit_add_circle` is equal to the integral over an interval (t, t + 1] in `ℝ` of its lift to `ℝ`. -/ -protected lemma lintegral_preimage (t : ℝ) {f : unit_add_circle → ℝ≥0∞} (hf : measurable f) : +protected lemma lintegral_preimage (t : ℝ) (f : unit_add_circle → ℝ≥0∞) : ∫⁻ a in Ioc t (t + 1), f a = ∫⁻ b : unit_add_circle, f b := -add_circle.lintegral_preimage 1 t hf +add_circle.lintegral_preimage 1 t f variables {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] [complete_space E] /-- The integral of an almost-everywhere strongly measurable function over `unit_add_circle` is equal to the integral over an interval (t, t + 1] in `ℝ` of its lift to `ℝ`. -/ -protected lemma integral_preimage (t : ℝ) {f : unit_add_circle → E} - (hf : ae_strongly_measurable f volume) : +protected lemma integral_preimage (t : ℝ) (f : unit_add_circle → E) : ∫ a in Ioc t (t + 1), f a = ∫ b : unit_add_circle, f b := -add_circle.integral_preimage 1 t hf +add_circle.integral_preimage 1 t f /-- The integral of an almost-everywhere strongly measurable function over `unit_add_circle` is equal to the integral over an interval (t, t + 1] in `ℝ` of its lift to `ℝ`. -/ -protected lemma interval_integral_preimage (t : ℝ) {f : unit_add_circle → E} - (hf : ae_strongly_measurable f volume) : +protected lemma interval_integral_preimage (t : ℝ) (f : unit_add_circle → E) : ∫ a in t..(t + 1), f a = ∫ b : unit_add_circle, f b := -add_circle.interval_integral_preimage 1 t hf +add_circle.interval_integral_preimage 1 t f end unit_add_circle diff --git a/src/topology/continuous_function/compact.lean b/src/topology/continuous_function/compact.lean index a9e1834733f46..159e697800626 100644 --- a/src/topology/continuous_function/compact.lean +++ b/src/topology/continuous_function/compact.lean @@ -238,6 +238,13 @@ def linear_isometry_bounded_of_compact : norm_map' := λ f, rfl, .. add_equiv_bounded_of_compact α E } +variables {α E} -- to match bounded_continuous_function.eval_clm + +/-- The evaluation at a point, as a continuous linear map from `C(α, 𝕜)` to `𝕜`. -/ +def eval_clm (x : α) : C(α, E) →L[𝕜] E := + (eval_clm 𝕜 x).comp + ((linear_isometry_bounded_of_compact α E 𝕜).to_linear_isometry).to_continuous_linear_map + end -- this lemma and the next are the analogues of those autogenerated by `@[simps]` for From 40acfb6aa7516ffe6f91136691df012a64683390 Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Mon, 9 Jan 2023 07:05:23 +0000 Subject: [PATCH 0181/1291] chore(algebra/module/linear_map): move `map_sum` lemmas (#18106) The lemmas `linear_map.map_sum` and `linear_equiv.map_sum` exist only for dot notation convenience, since `linear_map`s are of `add_monoid_hom_class` and therefore the generic lemma `map_sum` can replace them. Since these lemmas are the only reason that finsets are imported to `algebra/module/linear_map`, I propose deleting that import and moving the lemmas to `linear_algebra/basic`. This should open several files for porting. --- src/algebra/module/equiv.lean | 4 ---- src/algebra/module/linear_map.lean | 9 +++------ src/algebra/module/submodule/basic.lean | 2 +- src/linear_algebra/basic.lean | 7 +++++++ 4 files changed, 11 insertions(+), 11 deletions(-) diff --git a/src/algebra/module/equiv.lean b/src/algebra/module/equiv.lean index 7808fe6759225..a9fc1ee36fbae 100644 --- a/src/algebra/module/equiv.lean +++ b/src/algebra/module/equiv.lean @@ -34,7 +34,6 @@ linear equiv, linear equivalences, linear isomorphism, linear isomorphic -/ open function -open_locale big_operators universes u u' v w x y z variables {R : Type*} {R₁ : Type*} {R₂ : Type*} {R₃ : Type*} @@ -375,9 +374,6 @@ theorem map_smul (e : N₁ ≃ₗ[R₁] N₂) (c : R₁) (x : N₁) : e (c • x) = c • e x := map_smulₛₗ e c x omit module_N₁ module_N₂ -@[simp] lemma map_sum {s : finset ι} (u : ι → M) : e (∑ i in s, u i) = ∑ i in s, e (u i) := -e.to_linear_map.map_sum - @[simp] theorem map_eq_zero_iff {x : M} : e x = 0 ↔ x = 0 := e.to_add_equiv.map_eq_zero_iff theorem map_ne_zero_iff {x : M} : e x ≠ 0 ↔ x ≠ 0 := diff --git a/src/algebra/module/linear_map.lean b/src/algebra/module/linear_map.lean index affad643cf562..ef23231b5c315 100644 --- a/src/algebra/module/linear_map.lean +++ b/src/algebra/module/linear_map.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Nathaniel Thomas, Jeremy Avigad, Johannes Hölzl, Mario Carneiro, Anne Baanen, Frédéric Dupuis, Heather Macbeth -/ -import algebra.big_operators.basic import algebra.hom.group_action import algebra.module.pi import algebra.star.basic @@ -53,8 +52,10 @@ To ensure that composition works smoothly for semilinear maps, we use the typecl linear map -/ +assert_not_exists submonoid +assert_not_exists finset + open function -open_locale big_operators universes u u' v w x y z variables {R : Type*} {R₁ : Type*} {R₂ : Type*} {R₃ : Type*} @@ -365,10 +366,6 @@ end restrict_scalars variable {R} -@[simp] lemma map_sum {ι} {t : finset ι} {g : ι → M} : - f (∑ i in t, g i) = (∑ i in t, f (g i)) := -f.to_add_monoid_hom.map_sum _ _ - theorem to_add_monoid_hom_injective : function.injective (to_add_monoid_hom : (M →ₛₗ[σ] M₃) → (M →+ M₃)) := λ f g h, ext $ add_monoid_hom.congr_fun h diff --git a/src/algebra/module/submodule/basic.lean b/src/algebra/module/submodule/basic.lean index edee030082de4..1ffe24a81fd9e 100644 --- a/src/algebra/module/submodule/basic.lean +++ b/src/algebra/module/submodule/basic.lean @@ -247,7 +247,7 @@ lemma injective_subtype : injective p.subtype := subtype.coe_injective /-- Note the `add_submonoid` version of this lemma is called `add_submonoid.coe_finset_sum`. -/ @[simp] lemma coe_sum (x : ι → p) (s : finset ι) : ↑(∑ i in s, x i) = ∑ i in s, (x i : M) := -p.subtype.map_sum +map_sum p.subtype _ _ section restrict_scalars variables (S) [semiring S] [module S M] [module R M] [has_smul S R] [is_scalar_tower S R M] diff --git a/src/linear_algebra/basic.lean b/src/linear_algebra/basic.lean index 5dec3792dc113..379751730d4b5 100644 --- a/src/linear_algebra/basic.lean +++ b/src/linear_algebra/basic.lean @@ -146,6 +146,10 @@ variables [ring_hom_comp_triple σ₁₃ σ₃₄ σ₁₄] [ring_hom_comp_tripl variables (f : M →ₛₗ[σ₁₂] M₂) (g : M₂ →ₛₗ[σ₂₃] M₃) include R R₂ +@[simp] lemma map_sum {ι : Type*} {t : finset ι} {g : ι → M} : + f (∑ i in t, g i) = (∑ i in t, f (g i)) := +f.to_add_monoid_hom.map_sum _ _ + theorem comp_assoc (h : M₃ →ₛₗ[σ₃₄] M₄) : ((h.comp g : M₂ →ₛₗ[σ₂₄] M₄).comp f : M →ₛₗ[σ₁₄] M₄) = h.comp (g.comp f : M →ₛₗ[σ₁₃] M₃) := rfl @@ -1514,6 +1518,9 @@ variables {σ₁₂ : R →+* R₂} {σ₂₁ : R₂ →+* R} variables {re₁₂ : ring_hom_inv_pair σ₁₂ σ₂₁} {re₂₁ : ring_hom_inv_pair σ₂₁ σ₁₂} variables (e e' : M ≃ₛₗ[σ₁₂] M₂) +@[simp] lemma map_sum {s : finset ι} (u : ι → M) : e (∑ i in s, u i) = ∑ i in s, e (u i) := +e.to_linear_map.map_sum + lemma map_eq_comap {p : submodule R M} : (p.map (e : M →ₛₗ[σ₁₂] M₂) : submodule R₂ M₂) = p.comap (e.symm : M₂ →ₛₗ[σ₂₁] M) := set_like.coe_injective $ by simp [e.image_eq_preimage] From 5709b0d8725255e76f47debca6400c07b5c2d8e6 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 9 Jan 2023 10:12:26 +0000 Subject: [PATCH 0182/1291] feat(order/complete_lattice): missing API lemmas about the prod instance (#18029) This adds lemmas about how `fst`, `snd`, and `swap` interact with `Sup` and `Inf`. --- src/order/complete_lattice.lean | 56 +++++++++++++++++++++++++++++++++ 1 file changed, 56 insertions(+) diff --git a/src/order/complete_lattice.lean b/src/order/complete_lattice.lean index a1cfade619ff4..32a017dad7375 100644 --- a/src/order/complete_lattice.lean +++ b/src/order/complete_lattice.lean @@ -1295,6 +1295,54 @@ instance [has_Sup α] [has_Sup β] : has_Sup (α × β) := instance [has_Inf α] [has_Inf β] : has_Inf (α × β) := ⟨λ s, (Inf (prod.fst '' s), Inf (prod.snd '' s))⟩ +variables {α β} + +lemma fst_Inf [has_Inf α] [has_Inf β] (s : set (α × β)) : (Inf s).fst = Inf (prod.fst '' s) := rfl + +lemma snd_Inf [has_Inf α] [has_Inf β] (s : set (α × β)) : (Inf s).snd = Inf (prod.snd '' s) := rfl + +lemma swap_Inf [has_Inf α] [has_Inf β] (s : set (α × β)) : (Inf s).swap = Inf (prod.swap '' s) := +ext + (congr_arg Inf $ image_comp prod.fst swap s : _) + (congr_arg Inf $ image_comp prod.snd swap s : _) + +lemma fst_Sup [has_Sup α] [has_Sup β] (s : set (α × β)) : (Sup s).fst = Sup (prod.fst '' s) := rfl + +lemma snd_Sup [has_Sup α] [has_Sup β] (s : set (α × β)) : (Sup s).snd = Sup (prod.snd '' s) := rfl + +lemma swap_Sup [has_Sup α] [has_Sup β] (s : set (α × β)) : (Sup s).swap = Sup (prod.swap '' s) := +ext + (congr_arg Sup $ image_comp prod.fst swap s : _) + (congr_arg Sup $ image_comp prod.snd swap s : _) + +lemma fst_infi [has_Inf α] [has_Inf β] (f : ι → α × β) : (infi f).fst = ⨅ i, (f i).fst := +congr_arg Inf (range_comp _ _).symm + +lemma snd_infi [has_Inf α] [has_Inf β] (f : ι → α × β) : (infi f).snd = ⨅ i, (f i).snd := +congr_arg Inf (range_comp _ _).symm + +lemma swap_infi [has_Inf α] [has_Inf β] (f : ι → α × β) : (infi f).swap = ⨅ i, (f i).swap := +by simp_rw [infi, swap_Inf, range_comp] + +lemma infi_mk [has_Inf α] [has_Inf β] (f : ι → α) (g : ι → β) : + (⨅ i, (f i, g i)) = (⨅ i, f i, ⨅ i, g i) := +congr_arg2 prod.mk (fst_infi _) (snd_infi _) + +lemma fst_supr [has_Sup α] [has_Sup β] (f : ι → α × β) : (supr f).fst = ⨆ i, (f i).fst := +congr_arg Sup (range_comp _ _).symm + +lemma snd_supr [has_Sup α] [has_Sup β] (f : ι → α × β) : (supr f).snd = ⨆ i, (f i).snd := +congr_arg Sup (range_comp _ _).symm + +lemma swap_supr [has_Sup α] [has_Sup β] (f : ι → α × β) : (supr f).swap = ⨆ i, (f i).swap := +by simp_rw [supr, swap_Sup, range_comp] + +lemma supr_mk [has_Sup α] [has_Sup β] (f : ι → α) (g : ι → β) : + (⨆ i, (f i, g i)) = (⨆ i, f i, ⨆ i, g i) := +congr_arg2 prod.mk (fst_supr _) (snd_supr _) + +variables (α β) + instance [complete_lattice α] [complete_lattice β] : complete_lattice (α × β) := { le_Sup := λ s p hab, ⟨le_Sup $ mem_image_of_mem _ hab, le_Sup $ mem_image_of_mem _ hab⟩, Sup_le := λ s p h, @@ -1311,6 +1359,14 @@ instance [complete_lattice α] [complete_lattice β] : complete_lattice (α × end prod +lemma Inf_prod [has_Inf α] [has_Inf β] {s : set α} {t : set β} (hs : s.nonempty) (ht : t.nonempty) : + Inf (s ×ˢ t) = (Inf s, Inf t) := +congr_arg2 prod.mk (congr_arg Inf $ fst_image_prod _ ht) (congr_arg Inf $ snd_image_prod hs _) + +lemma Sup_prod [has_Sup α] [has_Sup β] {s : set α} {t : set β} (hs : s.nonempty) (ht : t.nonempty) : + Sup (s ×ˢ t) = (Sup s, Sup t) := +congr_arg2 prod.mk (congr_arg Sup $ fst_image_prod _ ht) (congr_arg Sup $ snd_image_prod hs _) + section complete_lattice variables [complete_lattice α] {a : α} {s : set α} From ba2d8fb3eea8a63fe2824547b53548c71e973a01 Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Mon, 9 Jan 2023 10:12:27 +0000 Subject: [PATCH 0183/1291] feat(ring_theory/ideal/operations): add span_singleton_mul_left lemmas (#18056) --- src/algebra/module/pid.lean | 6 +++--- src/ring_theory/ideal/basic.lean | 5 +++-- src/ring_theory/ideal/operations.lean | 25 +++++++++++++++++++++++++ 3 files changed, 31 insertions(+), 5 deletions(-) diff --git a/src/algebra/module/pid.lean b/src/algebra/module/pid.lean index 8c50e5b3ace4f..4b3a7fc1c6440 100644 --- a/src/algebra/module/pid.lean +++ b/src/algebra/module/pid.lean @@ -131,10 +131,10 @@ begin { rw [← quotient.mk_eq_zero, mk_smul, f1.some_spec, ← f.map_smul], convert f.map_zero, change _ • submodule.quotient.mk _ = _, rw [← mk_smul, quotient.mk_eq_zero, algebra.id.smul_eq_mul, mul_one], - exact mem_span_singleton_self _ }, + exact submodule.mem_span_singleton_self _ }, obtain ⟨a, ha⟩ := p_pow_smul_lift hp hM hz this, refine ⟨f1.some - a • z, by rw [smul_sub, sub_eq_zero, ha], _⟩, - rw [mk_sub, mk_smul, (quotient.mk_eq_zero _).mpr $ mem_span_singleton_self _, + rw [mk_sub, mk_smul, (quotient.mk_eq_zero _).mpr $ submodule.mem_span_singleton_self _, smul_zero, sub_zero, f1.some_spec] end @@ -194,7 +194,7 @@ begin rw [submodule.map_span, submodule.map_top, range_mkq] at hs', simp only [mkq_apply] at hs', simp only [s'], rw [set.range_comp (_ ∘ s), fin.range_succ_above], rw [← set.range_comp, ← set.insert_image_compl_eq_range _ j, function.comp_apply, - (quotient.mk_eq_zero _).mpr (mem_span_singleton_self _), span_insert_zero] at hs', + (quotient.mk_eq_zero _).mpr (submodule.mem_span_singleton_self _), span_insert_zero] at hs', exact hs' } } end end p_torsion diff --git a/src/ring_theory/ideal/basic.lean b/src/ring_theory/ideal/basic.lean index 4039a1588777d..bbc48e8f40948 100644 --- a/src/ring_theory/ideal/basic.lean +++ b/src/ring_theory/ideal/basic.lean @@ -374,10 +374,11 @@ variables [comm_semiring α] (I : ideal α) theorem mul_unit_mem_iff_mem {x y : α} (hy : is_unit y) : x * y ∈ I ↔ x ∈ I := mul_comm y x ▸ unit_mul_mem_iff_mem I hy -lemma mem_span_singleton {x y : α} : - x ∈ span ({y} : set α) ↔ y ∣ x := +lemma mem_span_singleton {x y : α} : x ∈ span ({y} : set α) ↔ y ∣ x := mem_span_singleton'.trans $ exists_congr $ λ _, by rw [eq_comm, mul_comm] +lemma mem_span_singleton_self (x : α) : x ∈ span ({x} : set α) := mem_span_singleton.mpr dvd_rfl + lemma span_singleton_le_span_singleton {x y : α} : span ({x} : set α) ≤ span ({y} : set α) ↔ y ∣ x := span_le.trans $ singleton_subset_iff.trans mem_span_singleton diff --git a/src/ring_theory/ideal/operations.lean b/src/ring_theory/ideal/operations.lean index 34eef05efe7f2..2c176382b96b5 100644 --- a/src/ring_theory/ideal/operations.lean +++ b/src/ring_theory/ideal/operations.lean @@ -443,6 +443,31 @@ lemma span_singleton_mul_le_span_singleton_mul {x y : R} {I J : ideal R} : span {x} * I ≤ span {y} * J ↔ ∀ zI ∈ I, ∃ zJ ∈ J, x * zI = y * zJ := by simp only [span_singleton_mul_le_iff, mem_span_singleton_mul, eq_comm] +lemma span_singleton_mul_right_mono [is_domain R] {x : R} (hx : x ≠ 0) : + span {x} * I ≤ span {x} * J ↔ I ≤ J := +by simp_rw [span_singleton_mul_le_span_singleton_mul, mul_right_inj' hx, exists_prop, + exists_eq_right', set_like.le_def] + +lemma span_singleton_mul_left_mono [is_domain R] {x : R} (hx : x ≠ 0) : + I * span {x} ≤ J * span {x} ↔ I ≤ J := +by simpa only [mul_comm I, mul_comm J] using span_singleton_mul_right_mono hx + +lemma span_singleton_mul_right_inj [is_domain R] {x : R} (hx : x ≠ 0) : + span {x} * I = span {x} * J ↔ I = J := +by simp only [le_antisymm_iff, span_singleton_mul_right_mono hx] + +lemma span_singleton_mul_left_inj [is_domain R] {x : R} (hx : x ≠ 0) : + I * span {x} = J * span {x} ↔ I = J := +by simp only [le_antisymm_iff, span_singleton_mul_left_mono hx] + +lemma span_singleton_mul_right_injective [is_domain R] {x : R} (hx : x ≠ 0) : + function.injective ((*) (span {x} : ideal R)) := +λ _ _, (span_singleton_mul_right_inj hx).mp + +lemma span_singleton_mul_left_injective [is_domain R] {x : R} (hx : x ≠ 0) : + function.injective (λ I : ideal R, I * span {x}) := +λ _ _, (span_singleton_mul_left_inj hx).mp + lemma eq_span_singleton_mul {x : R} (I J : ideal R) : I = span {x} * J ↔ ((∀ zI ∈ I, ∃ zJ ∈ J, x * zJ = zI) ∧ (∀ z ∈ J, x * z ∈ I)) := by simp only [le_antisymm_iff, le_span_singleton_mul_iff, span_singleton_mul_le_iff] From 51a845f04bd140e921c29a79e7a31e0ae2c25d1b Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Mon, 9 Jan 2023 13:44:18 +0000 Subject: [PATCH 0184/1291] feat(ring_theory/ideal/*): add more ideal.span lemmas (#18079) --- src/ring_theory/ideal/basic.lean | 4 ++++ src/ring_theory/ideal/operations.lean | 4 ++++ 2 files changed, 8 insertions(+) diff --git a/src/ring_theory/ideal/basic.lean b/src/ring_theory/ideal/basic.lean index bbc48e8f40948..9769b03b72100 100644 --- a/src/ring_theory/ideal/basic.lean +++ b/src/ring_theory/ideal/basic.lean @@ -527,6 +527,10 @@ protected lemma sub_mem : a ∈ I → b ∈ I → a - b ∈ I := sub_mem lemma mem_span_insert' {s : set α} {x y} : x ∈ span (insert y s) ↔ ∃a, x + a * y ∈ span s := submodule.mem_span_insert' +@[simp] lemma span_singleton_neg (x : α) : (span {-x} : ideal α) = span {x} := +by { ext, simp only [mem_span_singleton'], + exact ⟨λ ⟨y, h⟩, ⟨-y, h ▸ neg_mul_comm y x⟩, λ ⟨y, h⟩, ⟨-y, h ▸ neg_mul_neg y x⟩⟩ } + end ideal end ring diff --git a/src/ring_theory/ideal/operations.lean b/src/ring_theory/ideal/operations.lean index 2c176382b96b5..35bd9658d727f 100644 --- a/src/ring_theory/ideal/operations.lean +++ b/src/ring_theory/ideal/operations.lean @@ -649,6 +649,10 @@ lemma prod_eq_bot {R : Type*} [comm_ring R] [is_domain R] {s : multiset (ideal R)} : s.prod = ⊥ ↔ ∃ I ∈ s, I = ⊥ := prod_zero_iff_exists_zero +lemma span_pair_mul_span_pair (w x y z : R) : + (span {w, x} : ideal R) * span {y, z} = span {w * y, w * z, x * y, x * z} := +by simp_rw [span_insert, sup_mul, mul_sup, span_singleton_mul_span_singleton, sup_assoc] + /-- The radical of an ideal `I` consists of the elements `r` such that `r^n ∈ I` for some `n`. -/ def radical (I : ideal R) : ideal R := { carrier := { r | ∃ n : ℕ, r ^ n ∈ I }, From 02c4026cbe3cd2122eb8ff196c80f24441037002 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 9 Jan 2023 13:44:19 +0000 Subject: [PATCH 0185/1291] fix(algebra/group_power/lemmas): fix the name of sub_one_zsmul (#18109) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The old name was `zsmul_sub_one`, which didn't makes sense for the lemma statement ```lean ∀ {G : Type u_2} [_inst_1 : add_group G] (a : G) (n : ℤ), (n - 1) • a = n • a + -a ``` --- src/algebra/group_power/lemmas.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/algebra/group_power/lemmas.lean b/src/algebra/group_power/lemmas.lean index 39cb2da938b71..7b6b520351d17 100644 --- a/src/algebra/group_power/lemmas.lean +++ b/src/algebra/group_power/lemmas.lean @@ -148,7 +148,7 @@ lemma zpow_add_one (a : G) : ∀ n : ℤ, a ^ (n + 1) = a ^ n * a exact zpow_neg_succ_of_nat _ _ end -@[to_additive zsmul_sub_one] +@[to_additive sub_one_zsmul] lemma zpow_sub_one (a : G) (n : ℤ) : a ^ (n - 1) = a ^ n * a⁻¹ := calc a ^ (n - 1) = a ^ (n - 1) * a * a⁻¹ : (mul_inv_cancel_right _ _).symm ... = a^n * a⁻¹ : by rw [← zpow_add_one, sub_add_cancel] From e1bccd6e40ae78370f01659715d3c948716e3b7e Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 9 Jan 2023 16:42:07 +0000 Subject: [PATCH 0186/1291] chore(data/int/order/basic): Fix lemma name (#17967) `int.coe_nat_abs` wasn't referring to the correct lemma. Change the name, add the lemma it should correspond to and tag both with `simp` and `norm_cast`. Co-authored-by: Eric Wieser --- archive/imo/imo2013_q5.lean | 2 +- src/algebra/euclidean_domain/instances.lean | 2 +- src/algebra/group_power/lemmas.lean | 7 +++--- src/analysis/normed/field/basic.lean | 2 +- src/data/int/basic.lean | 4 +-- src/data/int/dvd/basic.lean | 4 +-- src/data/int/order/basic.lean | 9 +++++-- src/data/rat/floor.lean | 3 +-- src/data/sign.lean | 2 +- src/dynamics/ergodic/add_circle.lean | 2 +- src/number_theory/zsqrtd/gaussian_int.lean | 6 ++--- src/ring_theory/int/basic.lean | 28 ++++++++++----------- src/testing/slim_check/sampleable.lean | 2 +- 13 files changed, 37 insertions(+), 36 deletions(-) diff --git a/archive/imo/imo2013_q5.lean b/archive/imo/imo2013_q5.lean index 261f9526df2d7..53cb83b0df515 100644 --- a/archive/imo/imo2013_q5.lean +++ b/archive/imo/imo2013_q5.lean @@ -103,7 +103,7 @@ begin ... = ((q.num.nat_abs : ℤ) : ℝ) : congr_arg coe (int.nat_abs_of_nonneg num_pos.le).symm ... ≤ f q.num.nat_abs : H4 q.num.nat_abs (int.nat_abs_pos_of_ne_zero num_pos.ne') - ... = f q.num : by { rw ←int.nat_abs_of_nonneg num_pos.le, norm_cast } + ... = f q.num : by rw [nat.cast_nat_abs, abs_of_nonneg num_pos.le] ... = f (q * q.denom) : by rw ←rat.mul_denom_eq_num ... ≤ f q * f q.denom : H1 q q.denom hq (nat.cast_pos.mpr q.pos), have h_f_denom_pos := diff --git a/src/algebra/euclidean_domain/instances.lean b/src/algebra/euclidean_domain/instances.lean index 59e62167a406d..994324f38030d 100644 --- a/src/algebra/euclidean_domain/instances.lean +++ b/src/algebra/euclidean_domain/instances.lean @@ -33,7 +33,7 @@ instance int.euclidean_domain : euclidean_domain ℤ := r := λ a b, a.nat_abs < b.nat_abs, r_well_founded := measure_wf (λ a, int.nat_abs a), remainder_lt := λ a b b0, int.coe_nat_lt.1 $ - by { rw [int.nat_abs_of_nonneg (int.mod_nonneg _ b0), ← int.abs_eq_nat_abs], + by { rw [int.nat_abs_of_nonneg (int.mod_nonneg _ b0), int.coe_nat_abs], exact int.mod_lt _ b0 }, mul_left_not_lt := λ a b b0, not_lt_of_ge $ by {rw [← mul_one a.nat_abs, int.nat_abs_mul], diff --git a/src/algebra/group_power/lemmas.lean b/src/algebra/group_power/lemmas.lean index 7b6b520351d17..28338034995b0 100644 --- a/src/algebra/group_power/lemmas.lean +++ b/src/algebra/group_power/lemmas.lean @@ -300,9 +300,9 @@ lemma abs_zsmul (n : ℤ) (a : α) : |n • a| = |n| • |a| := begin obtain n0 | n0 := le_total 0 n, { lift n to ℕ using n0, - simp only [abs_nsmul, coe_nat_abs, coe_nat_zsmul] }, + simp only [abs_nsmul, abs_coe_nat, coe_nat_zsmul] }, { lift (- n) to ℕ using neg_nonneg.2 n0 with m h, - rw [← abs_neg (n • a), ← neg_zsmul, ← abs_neg n, ← h, coe_nat_zsmul, coe_nat_abs, + rw [← abs_neg (n • a), ← neg_zsmul, ← abs_neg n, ← h, coe_nat_zsmul, abs_coe_nat, coe_nat_zsmul], exact abs_nsmul m _ }, end @@ -524,8 +524,7 @@ end linear_ordered_ring namespace int -@[simp] lemma nat_abs_sq (x : ℤ) : (x.nat_abs ^ 2 : ℤ) = x ^ 2 := -by rw [sq, int.nat_abs_mul_self', sq] +lemma nat_abs_sq (x : ℤ) : (x.nat_abs ^ 2 : ℤ) = x ^ 2 := by rw [sq, int.nat_abs_mul_self', sq] alias nat_abs_sq ← nat_abs_pow_two diff --git a/src/analysis/normed/field/basic.lean b/src/analysis/normed/field/basic.lean index 8e327edcd3906..d7dd051897b82 100644 --- a/src/analysis/normed/field/basic.lean +++ b/src/analysis/normed/field/basic.lean @@ -699,7 +699,7 @@ lemma int.norm_eq_abs (n : ℤ) : ‖n‖ = |n| := rfl lemma nnreal.coe_nat_abs (n : ℤ) : (n.nat_abs : ℝ≥0) = ‖n‖₊ := nnreal.eq $ calc ((n.nat_abs : ℝ≥0) : ℝ) = (n.nat_abs : ℤ) : by simp only [int.cast_coe_nat, nnreal.coe_nat_cast] - ... = |n| : by simp only [← int.abs_eq_nat_abs, int.cast_abs] + ... = |n| : by simp only [int.coe_nat_abs, int.cast_abs] ... = ‖n‖ : rfl lemma int.abs_le_floor_nnreal_iff (z : ℤ) (c : ℝ≥0) : |z| ≤ ⌊c⌋₊ ↔ ‖z‖₊ ≤ c := diff --git a/src/data/int/basic.lean b/src/data/int/basic.lean index 52b27c99b41d9..433b7cb209ace 100644 --- a/src/data/int/basic.lean +++ b/src/data/int/basic.lean @@ -181,7 +181,7 @@ end variables {a b : ℤ} {n : ℕ} -attribute [simp] nat_abs nat_abs_of_nat nat_abs_zero nat_abs_one +attribute [simp] nat_abs_of_nat nat_abs_zero nat_abs_one theorem nat_abs_add_le (a b : ℤ) : nat_abs (a + b) ≤ nat_abs a + nat_abs b := begin @@ -212,7 +212,7 @@ lemma nat_abs_mul_nat_abs_eq {a b : ℤ} {c : ℕ} (h : a * b = (c : ℤ)) : a.nat_abs * b.nat_abs = c := by rw [← nat_abs_mul, h, nat_abs_of_nat] -@[simp] lemma nat_abs_mul_self' (a : ℤ) : (nat_abs a * nat_abs a : ℤ) = a * a := +lemma nat_abs_mul_self' (a : ℤ) : (nat_abs a * nat_abs a : ℤ) = a * a := by rw [← int.coe_nat_mul, nat_abs_mul_self] theorem neg_succ_of_nat_eq' (m : ℕ) : -[1+ m] = -m - 1 := diff --git a/src/data/int/dvd/basic.lean b/src/data/int/dvd/basic.lean index 034d16b7c309a..70b50fe2e31a9 100644 --- a/src/data/int/dvd/basic.lean +++ b/src/data/int/dvd/basic.lean @@ -27,10 +27,10 @@ namespace int λ ⟨k, e⟩, dvd.intro k $ by rw [e, int.coe_nat_mul]⟩ theorem coe_nat_dvd_left {n : ℕ} {z : ℤ} : (↑n : ℤ) ∣ z ↔ n ∣ z.nat_abs := -by rcases nat_abs_eq z with eq | eq; rw eq; simp [coe_nat_dvd] +by rcases nat_abs_eq z with eq | eq; rw eq; simp [←coe_nat_dvd] theorem coe_nat_dvd_right {n : ℕ} {z : ℤ} : z ∣ (↑n : ℤ) ↔ z.nat_abs ∣ n := -by rcases nat_abs_eq z with eq | eq; rw eq; simp [coe_nat_dvd] +by rcases nat_abs_eq z with eq | eq; rw eq; simp [←coe_nat_dvd] theorem le_of_dvd {a b : ℤ} (bpos : 0 < b) (H : a ∣ b) : a ≤ b := match a, b, eq_succ_of_zero_lt bpos, H with diff --git a/src/data/int/order/basic.lean b/src/data/int/order/basic.lean index de24db37ca9d1..94841ad2c1297 100644 --- a/src/data/int/order/basic.lean +++ b/src/data/int/order/basic.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad -/ import data.int.basic +import data.int.cast.basic import algebra.ring.divisibility import algebra.order.group.abs import algebra.order.ring.char_zero @@ -54,6 +55,11 @@ theorem abs_eq_nat_abs : ∀ a : ℤ, |a| = nat_abs a | (n : ℕ) := abs_of_nonneg $ coe_zero_le _ | -[1+ n] := abs_of_nonpos $ le_of_lt $ neg_succ_lt_zero _ +@[simp, norm_cast] lemma coe_nat_abs (n : ℤ) : (n.nat_abs : ℤ) = |n| := n.abs_eq_nat_abs.symm + +lemma _root_.nat.cast_nat_abs {α : Type*} [add_group_with_one α] (n : ℤ) : (n.nat_abs : α) = ↑|n| := +by rw [←int.coe_nat_abs, int.cast_coe_nat] + theorem nat_abs_abs (a : ℤ) : nat_abs (|a|) = nat_abs a := by rw [abs_eq_nat_abs]; refl @@ -68,8 +74,7 @@ lemma coe_nat_ne_zero_iff_pos {n : ℕ} : (n : ℤ) ≠ 0 ↔ 0 < n := ⟨λ h, nat.pos_of_ne_zero (coe_nat_ne_zero.1 h), λ h, (ne_of_lt (coe_nat_lt.2 h)).symm⟩ -theorem coe_nat_abs (n : ℕ) : |(n : ℤ)| = n := -abs_of_nonneg (coe_nat_nonneg n) +@[norm_cast] lemma abs_coe_nat (n : ℕ) : |(n : ℤ)| = n := abs_of_nonneg (coe_nat_nonneg n) /-! ### succ and pred -/ diff --git a/src/data/rat/floor.lean b/src/data/rat/floor.lean index b5096eae42c92..5428c18f5740e 100644 --- a/src/data/rat/floor.lean +++ b/src/data/rat/floor.lean @@ -126,9 +126,8 @@ begin have : ((q.denom - q.num * ⌊q_inv⌋ : ℚ) / q.num).num = q.denom - q.num * ⌊q_inv⌋, by { suffices : ((q.denom : ℤ) - q.num * ⌊q_inv⌋).nat_abs.coprime q.num.nat_abs, by exact_mod_cast (rat.num_div_eq_of_coprime q_num_pos this), - have : (q.num.nat_abs : ℚ) = (q.num : ℚ), by exact_mod_cast q_num_abs_eq_q_num, have tmp := nat.coprime_sub_mul_floor_rat_div_of_coprime q.cop.symm, - simpa only [this, q_num_abs_eq_q_num] using tmp }, + simpa only [nat.cast_nat_abs, abs_of_nonneg q_num_pos.le] using tmp }, rwa this }, -- to show the claim, start with the following inequality have q_inv_num_denom_ineq : q⁻¹.num - ⌊q⁻¹⌋ * q⁻¹.denom < q⁻¹.denom, by diff --git a/src/data/sign.lean b/src/data/sign.lean index 99bde85fc7a58..0e2013bcf4553 100644 --- a/src/data/sign.lean +++ b/src/data/sign.lean @@ -379,7 +379,7 @@ begin λ a, a.1, λ a, a.1.prop, _, _⟩, { simp [@sum_attach _ _ _ _ (λ a, (f a).nat_abs)] }, { intros x hx, - simp [sum_sigma, hx, ← int.sign_eq_sign, int.sign_mul_nat_abs, mul_comm ((f _).nat_abs : ℤ), + simp [sum_sigma, hx, ← int.sign_eq_sign, int.sign_mul_abs, mul_comm (|f _|), @sum_attach _ _ _ _ (λ a, ∑ j in range (f a).nat_abs, if a = x then (f a).sign else 0)] } end diff --git a/src/dynamics/ergodic/add_circle.lean b/src/dynamics/ergodic/add_circle.lean index 5a22777975888..a1c40c7ea404e 100644 --- a/src/dynamics/ergodic/add_circle.lean +++ b/src/dynamics/ergodic/add_circle.lean @@ -102,7 +102,7 @@ lemma ergodic_zsmul {n : ℤ} (hn : 1 < |n|) : ergodic (λ (y : add_circle T), n have hu₀ : ∀ j, add_order_of (u j) = n.nat_abs^j, { exact λ j, add_order_of_div_of_gcd_eq_one (pow_pos (pos_of_gt hn) j) (gcd_one_left _), }, have hnu : ∀ j, n^j • (u j) = 0 := λ j, by rw [← add_order_of_dvd_iff_zsmul_eq_zero, hu₀, - int.coe_nat_pow, ← int.abs_eq_nat_abs, ← abs_pow, abs_dvd], + int.coe_nat_pow, int.coe_nat_abs, ← abs_pow, abs_dvd], have hu₁ : ∀ j, ((u j) +ᵥ s : set _) =ᵐ[volume] s := λ j, by rw vadd_eq_self_of_preimage_zsmul_eq_self hs' (hnu j), have hu₂ : tendsto (λ j, add_order_of $ u j) at_top at_top, diff --git a/src/number_theory/zsqrtd/gaussian_int.lean b/src/number_theory/zsqrtd/gaussian_int.lean index 02bd567c03491..8aa9a9f782990 100644 --- a/src/number_theory/zsqrtd/gaussian_int.lean +++ b/src/number_theory/zsqrtd/gaussian_int.lean @@ -96,12 +96,12 @@ by rw [← @int.cast_inj ℝ _ _ _]; simp lemma norm_pos {x : ℤ[i]} : 0 < norm x ↔ x ≠ 0 := by rw [lt_iff_le_and_ne, ne.def, eq_comm, norm_eq_zero]; simp [norm_nonneg] -lemma coe_nat_abs_norm (x : ℤ[i]) : (x.norm.nat_abs : ℤ) = x.norm := +lemma abs_coe_nat_norm (x : ℤ[i]) : (x.norm.nat_abs : ℤ) = x.norm := int.nat_abs_of_nonneg (norm_nonneg _) @[simp] lemma nat_cast_nat_abs_norm {α : Type*} [ring α] (x : ℤ[i]) : (x.norm.nat_abs : α) = x.norm := -by rw [← int.cast_coe_nat, coe_nat_abs_norm] +by rw [← int.cast_coe_nat, abs_coe_nat_norm] lemma nat_abs_norm_eq (x : ℤ[i]) : x.norm.nat_abs = x.re.nat_abs * x.re.nat_abs + x.im.nat_abs * x.im.nat_abs := @@ -169,7 +169,7 @@ lemma norm_le_norm_mul_left (x : ℤ[i]) {y : ℤ[i]} (hy : y ≠ 0) : (norm x).nat_abs ≤ (norm (x * y)).nat_abs := by rw [zsqrtd.norm_mul, int.nat_abs_mul]; exact le_mul_of_one_le_right (nat.zero_le _) - (int.coe_nat_le.1 (by rw [coe_nat_abs_norm]; exact int.add_one_le_of_lt (norm_pos.2 hy))) + (int.coe_nat_le.1 (by rw [abs_coe_nat_norm]; exact int.add_one_le_of_lt (norm_pos.2 hy))) instance : nontrivial ℤ[i] := ⟨⟨0, 1, dec_trivial⟩⟩ diff --git a/src/ring_theory/int/basic.lean b/src/ring_theory/int/basic.lean index 5b1ce2c369f36..cd86ca03afeb9 100644 --- a/src/ring_theory/int/basic.lean +++ b/src/ring_theory/int/basic.lean @@ -91,24 +91,22 @@ instance : normalization_monoid ℤ := lemma normalize_of_nonneg {z : ℤ} (h : 0 ≤ z) : normalize z = z := show z * ↑(ite _ _ _) = z, by rw [if_pos h, units.coe_one, mul_one] -lemma normalize_of_neg {z : ℤ} (h : z < 0) : normalize z = -z := -show z * ↑(ite _ _ _) = -z, -by rw [if_neg (not_le_of_gt h), units.coe_neg, units.coe_one, mul_neg_one] +lemma normalize_of_nonpos {z : ℤ} (h : z ≤ 0) : normalize z = -z := +begin + obtain rfl | h := h.eq_or_lt, + { simp }, + { change z * ↑(ite _ _ _) = -z, + rw [if_neg (not_le_of_gt h), units.coe_neg, units.coe_one, mul_neg_one] } +end lemma normalize_coe_nat (n : ℕ) : normalize (n : ℤ) = n := normalize_of_nonneg (coe_nat_le_coe_nat_of_le $ nat.zero_le n) -theorem coe_nat_abs_eq_normalize (z : ℤ) : (z.nat_abs : ℤ) = normalize z := -begin - by_cases 0 ≤ z, - { simp [nat_abs_of_nonneg h, normalize_of_nonneg h] }, - { simp [of_nat_nat_abs_of_nonpos (le_of_not_ge h), normalize_of_neg (lt_of_not_ge h)] } -end +lemma abs_eq_normalize (z : ℤ) : |z| = normalize z := +by cases le_total 0 z; simp [normalize_of_nonneg, normalize_of_nonpos, *] lemma nonneg_of_normalize_eq_self {z : ℤ} (hz : normalize z = z) : 0 ≤ z := -calc 0 ≤ (z.nat_abs : ℤ) : coe_zero_le _ -... = normalize z : coe_nat_abs_eq_normalize _ -... = z : hz +abs_eq_self.1 $ by rw [abs_eq_normalize, hz] lemma nonneg_iff_normalize_eq_self (z : ℤ) : normalize z = z ↔ 0 ≤ z := ⟨nonneg_of_normalize_eq_self, normalize_of_nonneg⟩ @@ -127,7 +125,7 @@ instance : gcd_monoid ℤ := gcd_dvd_right := assume a b, int.gcd_dvd_right _ _, dvd_gcd := assume a b c, dvd_gcd, gcd_mul_lcm := λ a b, by - { rw [← int.coe_nat_mul, gcd_mul_lcm, coe_nat_abs_eq_normalize], + { rw [← int.coe_nat_mul, gcd_mul_lcm, coe_nat_abs, abs_eq_normalize], exact normalize_associated (a * b) }, lcm_zero_left := assume a, coe_nat_eq_zero.2 $ nat.lcm_zero_left _, lcm_zero_right := assume a, coe_nat_eq_zero.2 $ nat.lcm_zero_right _} @@ -225,10 +223,10 @@ begin { refine (assume a, quotient.induction_on' a $ assume a, associates.mk_eq_mk_iff_associated.2 $ associated.symm $ ⟨norm_unit a, _⟩), show normalize a = int.nat_abs (normalize a), - rw [int.coe_nat_abs_eq_normalize, normalize_idem] }, + rw [int.coe_nat_abs, int.abs_eq_normalize, normalize_idem] }, { intro n, dsimp, - rw [←normalize_apply, ← int.coe_nat_abs_eq_normalize, int.nat_abs_of_nat, int.nat_abs_of_nat] } + rw [←normalize_apply, ←int.abs_eq_normalize, int.nat_abs_abs, int.nat_abs_of_nat] } end lemma int.prime.dvd_mul {m n : ℤ} {p : ℕ} diff --git a/src/testing/slim_check/sampleable.lean b/src/testing/slim_check/sampleable.lean index 3da36e6b25d90..c14c317264fde 100644 --- a/src/testing/slim_check/sampleable.lean +++ b/src/testing/slim_check/sampleable.lean @@ -371,7 +371,7 @@ begin rcases i with ⟨x,⟨y,hy⟩⟩; unfold_wf; dsimp [rat.mk_pnat], mono*, - { rw [← int.coe_nat_le, ← int.abs_eq_nat_abs, ← int.abs_eq_nat_abs], + { rw [← int.coe_nat_le, int.coe_nat_abs, int.coe_nat_abs], apply int.abs_div_le_abs }, { change _ - 1 ≤ y-1, apply tsub_le_tsub_right, From 0b9c21e26ab115c62d3069a33320f950c51297be Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Mon, 9 Jan 2023 16:42:08 +0000 Subject: [PATCH 0187/1291] feat(algebraic_geometry/elliptic_curve/weierstrass): define ideals of the coordinate ring associated to the X and Y coordinates (#18038) Also slightly recategorise the sections. --- .../elliptic_curve/weierstrass.lean | 104 ++++++++++++------ 1 file changed, 70 insertions(+), 34 deletions(-) diff --git a/src/algebraic_geometry/elliptic_curve/weierstrass.lean b/src/algebraic_geometry/elliptic_curve/weierstrass.lean index ac7082cf315e8..5e1e64ed890f9 100644 --- a/src/algebraic_geometry/elliptic_curve/weierstrass.lean +++ b/src/algebraic_geometry/elliptic_curve/weierstrass.lean @@ -5,7 +5,7 @@ Authors: Kevin Buzzard, David Kurniadi Angdinata -/ import algebra.cubic_discriminant -import ring_theory.adjoin_root +import ring_theory.class_group import tactic.linear_combination /-! @@ -237,7 +237,7 @@ end torsion_polynomial section polynomial -/-! ### Weierstrass polynomials -/ +/-! ### Weierstrass polynomials and equations -/ open polynomial @@ -246,13 +246,43 @@ open_locale polynomial /-- The polynomial $W(X, Y) := Y^2 + a_1XY + a_3Y - (X^3 + a_2X^2 + a_4X + a_6)$ associated to a Weierstrass curve `W` over `R`. For ease of polynomial manipulation, this is represented as a term of type `R[X][X]`, where the inner variable represents $X$ and the outer variable represents $Y$. -/ -noncomputable def polynomial : R[X][X] := +protected noncomputable def polynomial : R[X][X] := X ^ 2 + C (C W.a₁ * X + C W.a₃) * X - C (X ^ 3 + C W.a₂ * X ^ 2 + C W.a₄ * X + C W.a₆) +lemma polynomial_eq : W.polynomial = cubic.to_poly + ⟨0, 1, cubic.to_poly ⟨0, 0, W.a₁, W.a₃⟩, cubic.to_poly ⟨-1, -W.a₂, -W.a₄, -W.a₆⟩⟩ := +by { simp only [weierstrass_curve.polynomial, cubic.to_poly, C_0, C_1, C_neg, C_add, C_mul], ring1 } + +lemma polynomial_ne_zero [nontrivial R] : W.polynomial ≠ 0 := +by { rw [polynomial_eq], exact cubic.ne_zero_of_b_ne_zero one_ne_zero } + +lemma degree_polynomial [nontrivial R] : W.polynomial.degree = 2 := +by { rw [polynomial_eq], exact cubic.degree_of_b_ne_zero' one_ne_zero } + +lemma nat_degree_polynomial [nontrivial R] : W.polynomial.nat_degree = 2 := +by { rw [polynomial_eq], exact cubic.nat_degree_of_b_ne_zero' one_ne_zero } + +lemma monic_polynomial : W.polynomial.monic := +by { nontriviality R, simpa only [polynomial_eq] using cubic.monic_of_b_eq_one' } + +lemma irreducible_polynomial [nontrivial R] [no_zero_divisors R] : irreducible W.polynomial := +begin + by_contra h, + rcases (W.monic_polynomial.not_irreducible_iff_exists_add_mul_eq_coeff W.nat_degree_polynomial).mp + h with ⟨f, g, h0, h1⟩, + simp only [polynomial_eq, cubic.coeff_eq_c, cubic.coeff_eq_d] at h0 h1, + apply_fun degree at h0 h1, + rw [cubic.degree_of_a_ne_zero' $ neg_ne_zero.mpr $ one_ne_zero' R, degree_mul] at h0, + apply (h1.symm.le.trans cubic.degree_of_b_eq_zero').not_lt, + rcases nat.with_bot.add_eq_three_iff.mp h0.symm with h | h | h | h, + any_goals { rw [degree_add_eq_left_of_degree_lt]; simp only [h]; dec_trivial }, + any_goals { rw [degree_add_eq_right_of_degree_lt]; simp only [h]; dec_trivial } +end + @[simp] lemma eval_polynomial (x y : R) : eval x (eval (C y) W.polynomial) = y ^ 2 + W.a₁ * x * y + W.a₃ * y - (x ^ 3 + W.a₂ * x ^ 2 + W.a₄ * x + W.a₆) := -by { simp only [polynomial], eval_simp, rw [add_mul, ← add_assoc] } +by { simp only [weierstrass_curve.polynomial], eval_simp, rw [add_mul, ← add_assoc] } @[simp] lemma eval_polynomial_zero : eval 0 (eval 0 W.polynomial) = -W.a₆ := by simp only [← C_0, eval_polynomial, zero_add, zero_sub, mul_zero, zero_pow (nat.zero_lt_succ _)] @@ -279,6 +309,8 @@ begin ring1 end +/-! ### Nonsingularity of Weierstrass curves -/ + /-- The partial derivative $W_X(X, Y)$ of $W(X, Y)$ with respect to $X$. -/ noncomputable def polynomial_X : R[X][X] := C (C W.a₁) * X - C (C 3 * X ^ 2 + C (2 * W.a₂) * X + C W.a₄) @@ -335,35 +367,7 @@ lemma nonsingular_of_Δ_ne_zero {x y : R} (h : W.equation x y) (hΔ : W.Δ ≠ 0 nonsingular_zero_of_Δ_ne_zero _ ((W.equation_iff_variable_change x y).mp h) $ by rwa [variable_change_Δ, inv_one, units.coe_one, one_pow, one_mul] -lemma polynomial_eq : W.polynomial = cubic.to_poly - ⟨0, 1, cubic.to_poly ⟨0, 0, W.a₁, W.a₃⟩, cubic.to_poly ⟨-1, -W.a₂, -W.a₄, -W.a₆⟩⟩ := -by { simp only [polynomial, cubic.to_poly, C_0, C_1, C_neg, C_add, C_mul], ring1 } - -lemma polynomial_ne_zero [nontrivial R] : W.polynomial ≠ 0 := -by { rw [polynomial_eq], exact cubic.ne_zero_of_b_ne_zero one_ne_zero } - -lemma polynomial_degree [nontrivial R] : W.polynomial.degree = 2 := -by { rw [polynomial_eq], exact cubic.degree_of_b_ne_zero' one_ne_zero } - -lemma polynomial_nat_degree [nontrivial R] : W.polynomial.nat_degree = 2 := -by { rw [polynomial_eq], exact cubic.nat_degree_of_b_ne_zero' one_ne_zero } - -lemma polynomial_monic : W.polynomial.monic := -by { nontriviality R, simpa only [polynomial_eq] using cubic.monic_of_b_eq_one' } - -lemma polynomial_irreducible [nontrivial R] [no_zero_divisors R] : irreducible W.polynomial := -begin - by_contra h, - rcases (W.polynomial_monic.not_irreducible_iff_exists_add_mul_eq_coeff W.polynomial_nat_degree).mp - h with ⟨f, g, h0, h1⟩, - simp only [polynomial_eq, cubic.coeff_eq_c, cubic.coeff_eq_d] at h0 h1, - apply_fun degree at h0 h1, - rw [cubic.degree_of_a_ne_zero' $ neg_ne_zero.mpr $ one_ne_zero' R, degree_mul] at h0, - apply (h1.symm.le.trans cubic.degree_of_b_eq_zero').not_lt, - rcases nat.with_bot.add_eq_three_iff.mp h0.symm with h | h | h | h, - any_goals { rw [degree_add_eq_left_of_degree_lt]; simp only [h]; dec_trivial }, - any_goals { rw [degree_add_eq_right_of_degree_lt]; simp only [h]; dec_trivial } -end +/-! ### The coordinate ring -/ /-- The coordinate ring $R[W] := R[X, Y] / \langle W(X, Y) \rangle$ of `W`. @@ -381,7 +385,7 @@ https://leanprover.zulipchat.com/#narrow/stream/116395-maths/topic/.E2.9C.94.20c instance [is_domain R] [normalized_gcd_monoid R] : is_domain W.coordinate_ring := (ideal.quotient.is_domain_iff_prime _).mpr $ by simpa only [ideal.span_singleton_prime W.polynomial_ne_zero, ← gcd_monoid.irreducible_iff_prime] - using W.polynomial_irreducible + using W.irreducible_polynomial instance coordinate_ring.is_domain_of_field {F : Type u} [field F] (W : weierstrass_curve F) : is_domain W.coordinate_ring := @@ -390,6 +394,38 @@ by { classical, apply_instance } /-- The function field $R(W) := \mathrm{Frac}(R[W])$ of `W`. -/ @[reducible] def function_field : Type u := fraction_ring W.coordinate_ring +variables (x : R) (y : R[X]) + +/-- The class of the element $X - x$ in $R[W]$ for some $x \in R$. -/ +@[simp] noncomputable def X_class : W.coordinate_ring := adjoin_root.mk W.polynomial $ C $ X - C x + +lemma X_class_ne_zero [nontrivial R] : W.X_class x ≠ 0 := +begin + intro hx, + cases ideal.mem_span_singleton'.mp (ideal.quotient.eq_zero_iff_mem.mp hx) with p hp, + apply_fun degree at hp, + rw [W.monic_polynomial.degree_mul, degree_polynomial, degree_C $ X_sub_C_ne_zero x] at hp, + cases p.degree; cases hp +end + +/-- The class of the element $Y - y(X)$ in $R[W]$ for some $y(X) \in R[X]$. -/ +@[simp] noncomputable def Y_class : W.coordinate_ring := adjoin_root.mk W.polynomial $ X - C y + +lemma Y_class_ne_zero [nontrivial R] : W.Y_class y ≠ 0 := +begin + intro hy, + cases ideal.mem_span_singleton'.mp (ideal.quotient.eq_zero_iff_mem.mp hy) with p hp, + apply_fun degree at hp, + rw [W.monic_polynomial.degree_mul, degree_polynomial, degree_X_sub_C] at hp, + cases p.degree; cases hp +end + +/-- The ideal $\langle X - x \rangle$ of $R[W]$ for some $x \in R$. -/ +@[simp] noncomputable def X_ideal : ideal W.coordinate_ring := ideal.span {W.X_class x} + +/-- The ideal $\langle Y - y(X) \rangle$ of $R[W]$ for some $y(X) \in R[X]$. -/ +@[simp] noncomputable def Y_ideal : ideal W.coordinate_ring := ideal.span {W.Y_class y} + end polynomial end weierstrass_curve From dd71334db81d0bd444af1ee339a29298bef40734 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 9 Jan 2023 16:42:09 +0000 Subject: [PATCH 0188/1291] feat(order/partial_sups): add csupr lemmas (#18053) --- src/order/partial_sups.lean | 49 ++++++++++++++++++++++++++++--------- 1 file changed, 37 insertions(+), 12 deletions(-) diff --git a/src/order/partial_sups.lean b/src/order/partial_sups.lean index fa4a0a5b400b3..41e97ca76bed3 100644 --- a/src/order/partial_sups.lean +++ b/src/order/partial_sups.lean @@ -5,13 +5,15 @@ Authors: Scott Morrison -/ import data.finset.lattice import order.hom.basic +import order.conditionally_complete_lattice.finset /-! # The monotone sequence of partial supremums of a sequence We define `partial_sups : (ℕ → α) → ℕ →o α` inductively. For `f : ℕ → α`, `partial_sups f` is the sequence `f 0 `, `f 0 ⊔ f 1`, `f 0 ⊔ f 1 ⊔ f 2`, ... The point of this definition is that -* it doesn't need a `⨆`, as opposed to `⨆ (i ≤ n), f i`. +* it doesn't need a `⨆`, as opposed to `⨆ (i ≤ n), f i` (which also means the wrong thing on + `conditionally_complete_lattice`s). * it doesn't need a `⊥`, as opposed to `(finset.range (n + 1)).sup f`. * it avoids needing to prove that `finset.range (n + 1)` is nonempty to use `finset.sup'`. @@ -67,6 +69,17 @@ begin { exact sup_le (ih (λ m p, w m (nat.le_succ_of_le p))) (w (n + 1) le_rfl) } end +@[simp] lemma bdd_above_range_partial_sups {f : ℕ → α} : + bdd_above (set.range (partial_sups f)) ↔ bdd_above (set.range f) := +begin + apply exists_congr (λ a, _), + split, + { rintros h b ⟨i, rfl⟩, + exact (le_partial_sups _ _).trans (h (set.mem_range_self i)) }, + { rintros h b ⟨i, rfl⟩, + exact (partial_sups_le _ _ _ $ λ _ _, h (set.mem_range_self _)), }, +end + lemma monotone.partial_sups_eq {f : ℕ → α} (hf : monotone f) : (partial_sups f : ℕ → α) = f := begin @@ -132,25 +145,37 @@ begin exact ⟨ih (nat.lt_of_succ_lt hmn), h hmn.ne⟩ } end +section conditionally_complete_lattice +variables [conditionally_complete_lattice α] + +lemma partial_sups_eq_csupr_Iic (f : ℕ → α) (n : ℕ) : partial_sups f n = ⨆ i : set.Iic n, f i := +begin + have : set.Iio (n + 1) = set.Iic n := set.ext (λ _, nat.lt_succ_iff), + rw [partial_sups_eq_sup'_range, finset.sup'_eq_cSup_image, finset.coe_range, + supr, set.range_comp, subtype.range_coe, this], +end + +@[simp] lemma csupr_partial_sups_eq {f : ℕ → α} (h : bdd_above (set.range f)) : + (⨆ n, partial_sups f n) = ⨆ n, f n := +begin + refine (csupr_le $ λ n, _).antisymm (csupr_mono _ $ le_partial_sups f), + { rw partial_sups_eq_csupr_Iic, + exact csupr_le (λ i, le_csupr h _), }, + { rwa bdd_above_range_partial_sups }, +end + +end conditionally_complete_lattice + section complete_lattice variables [complete_lattice α] lemma partial_sups_eq_bsupr (f : ℕ → α) (n : ℕ) : partial_sups f n = ⨆ (i ≤ n), f i := -begin - rw [partial_sups_eq_sup_range, finset.sup_eq_supr], - congr, - ext a, - exact supr_congr_Prop (by rw [finset.mem_range, nat.lt_succ_iff]) (λ _, rfl), -end +by simpa only [supr_subtype] using partial_sups_eq_csupr_Iic f n @[simp] lemma supr_partial_sups_eq (f : ℕ → α) : (⨆ n, partial_sups f n) = ⨆ n, f n := -begin - refine (supr_le $ λ n, _).antisymm (supr_mono $ le_partial_sups f), - rw partial_sups_eq_bsupr, - exact supr₂_le_supr _ _, -end +csupr_partial_sups_eq $ order_top.bdd_above _ lemma supr_le_supr_of_partial_sups_le_partial_sups {f g : ℕ → α} (h : partial_sups f ≤ partial_sups g) : From 7ea604785a41a0681eac70c5a82372493dbefc68 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 9 Jan 2023 19:56:46 +0000 Subject: [PATCH 0189/1291] feat(algebra/order/absolute_value): Absolute values are multiplicative ring norms (#16919) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Generalise `..._seminorm_class` to arbitrary codomains (rather than just `ℝ`). Instantiate `mul_ring_norm_class` for absolute values. --- src/algebra/order/absolute_value.lean | 6 + src/algebra/order/hom/basic.lean | 190 +++++++++++++++++++++++- src/analysis/normed/group/seminorm.lean | 97 +----------- src/analysis/normed/ring/seminorm.lean | 49 ++---- src/analysis/seminorm.lean | 2 +- 5 files changed, 212 insertions(+), 132 deletions(-) diff --git a/src/algebra/order/absolute_value.lean b/src/algebra/order/absolute_value.lean index adfde39ee76c6..906537205350e 100644 --- a/src/algebra/order/absolute_value.lean +++ b/src/algebra/order/absolute_value.lean @@ -175,6 +175,12 @@ by rw [← neg_sub, abv.map_neg] end ordered_comm_ring +instance {R S : Type*} [ring R] [ordered_comm_ring S] [nontrivial R] [is_domain S] : + mul_ring_norm_class (absolute_value R S) R S := +{ map_neg_eq_map := λ f, f.map_neg, + eq_zero_of_map_eq_zero := λ f a, f.eq_zero.1, + ..absolute_value.subadditive_hom_class, ..absolute_value.monoid_with_zero_hom_class } + section linear_ordered_ring variables {R S : Type*} [semiring R] [linear_ordered_ring S] (abv : absolute_value R S) diff --git a/src/algebra/order/hom/basic.lean b/src/algebra/order/hom/basic.lean index bc790b97dd7d9..26cdd079ec5f5 100644 --- a/src/algebra/order/hom/basic.lean +++ b/src/algebra/order/hom/basic.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ -import algebra.hom.group +import algebra.group_power.order /-! # Algebraic order homomorphism classes @@ -15,23 +15,67 @@ This file defines hom classes for common properties at the intersection of order ## Typeclasses +Basic typeclasses * `nonneg_hom_class`: Homs are nonnegative: `∀ f a, 0 ≤ f a` * `subadditive_hom_class`: Homs are subadditive: `∀ f a b, f (a + b) ≤ f a + f b` * `submultiplicative_hom_class`: Homs are submultiplicative: `∀ f a b, f (a * b) ≤ f a * f b` * `mul_le_add_hom_class`: `∀ f a b, f (a * b) ≤ f a + f b` * `nonarchimedean_hom_class`: `∀ a b, f (a + b) ≤ max (f a) (f b)` +Group norms +* `add_group_seminorm_class`: Homs are nonnegative, subadditive, even and preserve zero. +* `group_seminorm_class`: Homs are nonnegative, respect `f (a * b) ≤ f a + f b`, `f a⁻¹ = f a` and + preserve zero. +* `add_group_norm_class`: Homs are seminorms such that `f x = 0 → x = 0` for all `x`. +* `group_norm_class`: Homs are seminorms such that `f x = 0 → x = 1` for all `x`. + +Ring norms +* `ring_seminorm_class`: Homs are submultiplicative group norms. +* `ring_norm_class`: Homs are ring seminorms that are also additive group norms. +* `mul_ring_seminorm_class`: Homs are ring seminorms that are multiplicative. +* `mul_ring_norm_class`: Homs are ring norms that are multiplicative. + +## Notes + +Typeclasses for seminorms are defined here while types of seminorms are defined in +`analysis.normed.group.seminorm` and `analysis.normed.ring.seminorm` because absolute values are +multiplicative ring norms but outside of this use we only consider real-valued seminorms. + ## TODO Finitary versions of the current lemmas. -/ +/-- +Diamond inheritance cannot depend on `out_param`s in the following circumstances: + * there are three classes `top`, `middle`, `bottom` + * all of these classes have a parameter `(α : out_param _)` + * all of these classes have an instance parameter `[root α]` that depends on this `out_param` + * the `root` class has two child classes: `left` and `right`, these are siblings in the hierarchy + * the instance `bottom.to_middle` takes a `[left α]` parameter + * the instance `middle.to_top` takes a `[right α]` parameter + * there is a `leaf` class that inherits from both `left` and `right`. +In that case, given instances `bottom α` and `leaf α`, Lean cannot synthesize a `top α` instance, +even though the hypotheses of the instances `bottom.to_middle` and `middle.to_top` are satisfied. + +There are two workarounds: +* You could replace the bundled inheritance implemented by the instance `middle.to_top` with + unbundled inheritance implemented by adding a `[top α]` parameter to the `middle` class. This is + the preferred option since it is also more compatible with Lean 4, at the cost of being more work + to implement and more verbose to use. +* You could weaken the `bottom.to_middle` instance by making it depend on a subclass of + `middle.to_top`'s parameter, in this example replacing `[left α]` with `[leaf α]`. +-/ +library_note "out-param inheritance" + set_option old_structure_cmd true open function variables {ι F α β γ δ : Type*} +/-! ### Basics -/ + /-- `nonneg_hom_class F α β` states that `F` is a type of nonnegative morphisms. -/ class nonneg_hom_class (F : Type*) (α β : out_param $ Type*) [has_zero β] [has_le β] extends fun_like F α (λ _, β) := @@ -84,3 +128,147 @@ by simpa only [div_mul_div_cancel'] using map_mul_le_mul f (a / b) (b / c) lemma le_map_div_add_map_div [group α] [add_comm_semigroup β] [has_le β] [mul_le_add_hom_class F α β] (f : F) (a b c: α) : f (a / c) ≤ f (a / b) + f (b / c) := by simpa only [div_mul_div_cancel'] using map_mul_le_add f (a / b) (b / c) + +/-! ### Group (semi)norms -/ + +/-- `add_group_seminorm_class F α` states that `F` is a type of `β`-valued seminorms on the additive +group `α`. + +You should extend this class when you extend `add_group_seminorm`. -/ +class add_group_seminorm_class (F : Type*) (α β : out_param $ Type*) [add_group α] + [ordered_add_comm_monoid β] extends subadditive_hom_class F α β := +(map_zero (f : F) : f 0 = 0) +(map_neg_eq_map (f : F) (a : α) : f (-a) = f a) + +/-- `group_seminorm_class F α` states that `F` is a type of `β`-valued seminorms on the group `α`. + +You should extend this class when you extend `group_seminorm`. -/ +@[to_additive] +class group_seminorm_class (F : Type*) (α β : out_param $ Type*) [group α] + [ordered_add_comm_monoid β] extends mul_le_add_hom_class F α β := +(map_one_eq_zero (f : F) : f 1 = 0) +(map_inv_eq_map (f : F) (a : α) : f a⁻¹ = f a) + +/-- `add_group_norm_class F α` states that `F` is a type of `β`-valued norms on the additive group +`α`. + +You should extend this class when you extend `add_group_norm`. -/ +class add_group_norm_class (F : Type*) (α β : out_param $ Type*) [add_group α] + [ordered_add_comm_monoid β] extends add_group_seminorm_class F α β := +(eq_zero_of_map_eq_zero (f : F) {a : α} : f a = 0 → a = 0) + +/-- `group_norm_class F α` states that `F` is a type of `β`-valued norms on the group `α`. + +You should extend this class when you extend `group_norm`. -/ +@[to_additive] +class group_norm_class (F : Type*) (α β : out_param $ Type*) [group α] [ordered_add_comm_monoid β] + extends group_seminorm_class F α β := +(eq_one_of_map_eq_zero (f : F) {a : α} : f a = 0 → a = 1) + +export add_group_seminorm_class (map_neg_eq_map) +export group_seminorm_class (map_one_eq_zero map_inv_eq_map) +export add_group_norm_class (eq_zero_of_map_eq_zero) +export group_norm_class (eq_one_of_map_eq_zero) + +attribute [simp, to_additive map_zero] map_one_eq_zero +attribute [simp] map_neg_eq_map +attribute [simp, to_additive] map_inv_eq_map +attribute [to_additive] group_seminorm_class.to_mul_le_add_hom_class +attribute [to_additive] group_norm_class.to_group_seminorm_class + +@[priority 100] -- See note [lower instance priority] +instance add_group_seminorm_class.to_zero_hom_class [add_group α] [ordered_add_comm_monoid β] + [add_group_seminorm_class F α β] : + zero_hom_class F α β := +{ ..‹add_group_seminorm_class F α β› } + +section group_seminorm_class +variables [group α] [ordered_add_comm_monoid β] [group_seminorm_class F α β] (f : F) (x y : α) +include α β + +@[to_additive] lemma map_div_le_add : f (x / y) ≤ f x + f y := +by { rw [div_eq_mul_inv, ←map_inv_eq_map f y], exact map_mul_le_add _ _ _ } + +@[to_additive] lemma map_div_rev : f (x / y) = f (y / x) := by rw [←inv_div, map_inv_eq_map] + +@[to_additive] lemma le_map_add_map_div' : f x ≤ f y + f (y / x) := +by simpa only [add_comm, map_div_rev, div_mul_cancel'] using map_mul_le_add f (x / y) y + +end group_seminorm_class + +example [ordered_add_comm_group β] : ordered_add_comm_monoid β := infer_instance + +@[to_additive] lemma abs_sub_map_le_div [group α] [linear_ordered_add_comm_group β] + [group_seminorm_class F α β] (f : F) (x y : α) : |f x - f y| ≤ f (x / y) := +begin + rw [abs_sub_le_iff, sub_le_iff_le_add', sub_le_iff_le_add'], + exact ⟨le_map_add_map_div _ _ _, le_map_add_map_div' _ _ _⟩ +end + +@[to_additive, priority 100] -- See note [lower instance priority] +instance group_seminorm_class.to_nonneg_hom_class [group α] [linear_ordered_add_comm_monoid β] + [group_seminorm_class F α β] : + nonneg_hom_class F α β := +{ map_nonneg := λ f a, (nsmul_nonneg_iff two_ne_zero).1 $ + by { rw [two_nsmul, ←map_one_eq_zero f, ←div_self' a], exact map_div_le_add _ _ _ }, + ..‹group_seminorm_class F α β› } + +section group_norm_class +variables [group α] [ordered_add_comm_monoid β] [group_norm_class F α β] (f : F) {x : α} +include α β + +@[simp, to_additive] lemma map_eq_zero_iff_eq_one : f x = 0 ↔ x = 1 := +⟨eq_one_of_map_eq_zero _, by { rintro rfl, exact map_one_eq_zero _ }⟩ + +@[to_additive] lemma map_ne_zero_iff_ne_one : f x ≠ 0 ↔ x ≠ 1 := (map_eq_zero_iff_eq_one _).not + +end group_norm_class + +@[to_additive] lemma map_pos_of_ne_one [group α] [linear_ordered_add_comm_monoid β] + [group_norm_class F α β] (f : F) {x : α} (hx : x ≠ 1) : 0 < f x := +(map_nonneg _ _).lt_of_ne $ ((map_ne_zero_iff_ne_one _).2 hx).symm + +/-! ### Ring (semi)norms -/ + +/-- `ring_seminorm_class F α` states that `F` is a type of `β`-valued seminorms on the ring `α`. + +You should extend this class when you extend `ring_seminorm`. -/ +class ring_seminorm_class (F : Type*) (α β : out_param $ Type*) [non_unital_non_assoc_ring α] + [ordered_semiring β] extends add_group_seminorm_class F α β, submultiplicative_hom_class F α β + +/-- `ring_norm_class F α` states that `F` is a type of `β`-valued norms on the ring `α`. + +You should extend this class when you extend `ring_norm`. -/ +class ring_norm_class (F : Type*) (α β : out_param $ Type*) [non_unital_non_assoc_ring α] + [ordered_semiring β] extends ring_seminorm_class F α β, add_group_norm_class F α β + +/-- `mul_ring_seminorm_class F α` states that `F` is a type of `β`-valued multiplicative seminorms +on the ring `α`. + +You should extend this class when you extend `mul_ring_seminorm`. -/ +class mul_ring_seminorm_class (F : Type*) (α β : out_param $ Type*) [non_assoc_ring α] + [ordered_semiring β] extends add_group_seminorm_class F α β, monoid_with_zero_hom_class F α β + +/-- `mul_ring_norm_class F α` states that `F` is a type of `β`-valued multiplicative norms on the +ring `α`. + +You should extend this class when you extend `mul_ring_norm`. -/ +class mul_ring_norm_class (F : Type*) (α β : out_param $ Type*) [non_assoc_ring α] + [ordered_semiring β] extends mul_ring_seminorm_class F α β, add_group_norm_class F α β + +-- See note [out-param inheritance] +@[priority 100] -- See note [lower instance priority] +instance ring_seminorm_class.to_nonneg_hom_class [non_unital_non_assoc_ring α] + [linear_ordered_semiring β] [ring_seminorm_class F α β] : nonneg_hom_class F α β := +add_group_seminorm_class.to_nonneg_hom_class + +@[priority 100] -- See note [lower instance priority] +instance mul_ring_seminorm_class.to_ring_seminorm_class [non_assoc_ring α] [ordered_semiring β] + [mul_ring_seminorm_class F α β] : ring_seminorm_class F α β := +{ map_mul_le_mul := λ f a b, (map_mul _ _ _).le, + ..‹mul_ring_seminorm_class F α β› } + +@[priority 100] -- See note [lower instance priority] +instance mul_ring_norm_class.to_ring_norm_class [non_assoc_ring α] [ordered_semiring β] + [mul_ring_norm_class F α β] : ring_norm_class F α β := +{ ..‹mul_ring_norm_class F α β›, ..mul_ring_seminorm_class.to_ring_seminorm_class } diff --git a/src/analysis/normed/group/seminorm.lean b/src/analysis/normed/group/seminorm.lean index 9ee6ab03cfe60..c84c02e26aee6 100644 --- a/src/analysis/normed/group/seminorm.lean +++ b/src/analysis/normed/group/seminorm.lean @@ -21,6 +21,11 @@ is positive-semidefinite and subadditive. A norm further only maps zero to zero. * `add_group_norm`: A seminorm `f` such that `f x = 0 → x = 0` for all `x`. * `group_norm`: A seminorm `f` such that `f x = 0 → x = 1` for all `x`. +## Notes + +The corresponding hom classes are defined in `analysis.order.hom.basic` to be used by absolute +values. + ## References * [H. H. Schaefer, *Topological Vector Spaces*][schaefer1966] @@ -67,95 +72,7 @@ structure group_norm (G : Type*) [group G] extends group_seminorm G := attribute [nolint doc_blame] add_group_seminorm.to_zero_hom add_group_norm.to_add_group_seminorm group_norm.to_group_seminorm -/-- `add_group_seminorm_class F α` states that `F` is a type of seminorms on the additive group `α`. - -You should extend this class when you extend `add_group_seminorm`. -/ -class add_group_seminorm_class (F : Type*) (α : out_param $ Type*) [add_group α] - extends subadditive_hom_class F α ℝ := -(map_zero (f : F) : f 0 = 0) -(map_neg_eq_map (f : F) (a : α) : f (-a) = f a) - -/-- `group_seminorm_class F α` states that `F` is a type of seminorms on the group `α`. - -You should extend this class when you extend `group_seminorm`. -/ -@[to_additive] -class group_seminorm_class (F : Type*) (α : out_param $ Type*) [group α] - extends mul_le_add_hom_class F α ℝ := -(map_one_eq_zero (f : F) : f 1 = 0) -(map_inv_eq_map (f : F) (a : α) : f a⁻¹ = f a) - -/-- `add_group_norm_class F α` states that `F` is a type of norms on the additive group `α`. - -You should extend this class when you extend `add_group_norm`. -/ -class add_group_norm_class (F : Type*) (α : out_param $ Type*) [add_group α] - extends add_group_seminorm_class F α := -(eq_zero_of_map_eq_zero (f : F) {a : α} : f a = 0 → a = 0) - -/-- `group_norm_class F α` states that `F` is a type of norms on the group `α`. - -You should extend this class when you extend `group_norm`. -/ -@[to_additive] -class group_norm_class (F : Type*) (α : out_param $ Type*) [group α] - extends group_seminorm_class F α := -(eq_one_of_map_eq_zero (f : F) {a : α} : f a = 0 → a = 1) - -export add_group_seminorm_class (map_neg_eq_map) - group_seminorm_class (map_one_eq_zero map_inv_eq_map) - add_group_norm_class (eq_zero_of_map_eq_zero) - group_norm_class (eq_one_of_map_eq_zero) - -attribute [simp, to_additive map_zero] map_one_eq_zero -attribute [simp] map_neg_eq_map -attribute [simp, to_additive] map_inv_eq_map -attribute [to_additive] group_seminorm_class.to_mul_le_add_hom_class attribute [to_additive] group_norm.to_group_seminorm -attribute [to_additive] group_norm_class.to_group_seminorm_class - -@[priority 100] -- See note [lower instance priority] -instance add_group_seminorm_class.to_zero_hom_class [add_group E] [add_group_seminorm_class F E] : - zero_hom_class F E ℝ := -{ ..‹add_group_seminorm_class F E› } - -section group_seminorm_class -variables [group E] [group_seminorm_class F E] (f : F) (x y : E) -include E - -@[to_additive] lemma map_div_le_add : f (x / y) ≤ f x + f y := -by { rw [div_eq_mul_inv, ←map_inv_eq_map f y], exact map_mul_le_add _ _ _ } - -@[to_additive] lemma map_div_rev : f (x / y) = f (y / x) := by rw [←inv_div, map_inv_eq_map] - -@[to_additive] lemma le_map_add_map_div' : f x ≤ f y + f (y / x) := -by simpa only [add_comm, map_div_rev, div_mul_cancel'] using map_mul_le_add f (x / y) y - -@[to_additive] lemma abs_sub_map_le_div : |f x - f y| ≤ f (x / y) := -begin - rw [abs_sub_le_iff, sub_le_iff_le_add', sub_le_iff_le_add'], - exact ⟨le_map_add_map_div _ _ _, le_map_add_map_div' _ _ _⟩ -end - -end group_seminorm_class - -@[to_additive, priority 100] -- See note [lower instance priority] -instance group_seminorm_class.to_nonneg_hom_class [group E] [group_seminorm_class F E] : - nonneg_hom_class F E ℝ := -{ map_nonneg := λ f a, nonneg_of_mul_nonneg_right - (by { rw [two_mul, ←map_one_eq_zero f, ←div_self' a], exact map_div_le_add _ _ _ }) two_pos, - ..‹group_seminorm_class F E› } - -section group_norm_class -variables [group E] [group_norm_class F E] (f : F) {x : E} -include E - -@[to_additive] lemma map_pos_of_ne_one (hx : x ≠ 1) : 0 < f x := -(map_nonneg _ _).lt_of_ne $ λ h, hx $ eq_one_of_map_eq_zero _ h.symm - -@[simp, to_additive] lemma map_eq_zero_iff_eq_one : f x = 0 ↔ x = 1 := -⟨eq_one_of_map_eq_zero _, by { rintro rfl, exact map_one_eq_zero _ }⟩ - -@[to_additive] lemma map_ne_zero_iff_ne_one : f x ≠ 0 ↔ x ≠ 1 := (map_eq_zero_iff_eq_one _).not - -end group_norm_class /-! ### Seminorms -/ @@ -163,7 +80,7 @@ namespace group_seminorm section group variables [group E] [group F] [group G] {p q : group_seminorm E} -@[to_additive] instance group_seminorm_class : group_seminorm_class (group_seminorm E) E := +@[to_additive] instance group_seminorm_class : group_seminorm_class (group_seminorm E) E ℝ := { coe := λ f, f.to_fun, coe_injective' := λ f g h, by cases f; cases g; congr', map_one_eq_zero := λ f, f.map_one', @@ -401,7 +318,7 @@ namespace group_norm section group variables [group E] [group F] [group G] {p q : group_norm E} -@[to_additive] instance group_norm_class : group_norm_class (group_norm E) E := +@[to_additive] instance group_norm_class : group_norm_class (group_norm E) E ℝ := { coe := λ f, f.to_fun, coe_injective' := λ f g h, by cases f; cases g; congr', map_one_eq_zero := λ f, f.map_one', diff --git a/src/analysis/normed/ring/seminorm.lean b/src/analysis/normed/ring/seminorm.lean index 37ffa61e48c85..c42154bb990a3 100644 --- a/src/analysis/normed/ring/seminorm.lean +++ b/src/analysis/normed/ring/seminorm.lean @@ -22,6 +22,11 @@ For a ring `R`: multiplication. * `mul_ring_norm`: A multiplicative norm on a ring `R` is a ring norm that preserves multiplication. +## Notes + +The corresponding hom classes are defined in `analysis.order.hom.basic` to be used by absolute +values. + ## References * [S. Bosch, U. Güntzer, R. Remmert, *Non-Archimedean Analysis*][bosch-guntzer-remmert] @@ -62,49 +67,13 @@ attribute [nolint doc_blame] ring_seminorm.to_add_group_seminorm ring_norm.to_ad mul_ring_seminorm.to_monoid_with_zero_hom mul_ring_norm.to_add_group_norm mul_ring_norm.to_mul_ring_seminorm -/-- `ring_seminorm_class F α` states that `F` is a type of seminorms on the ring `α`. - -You should extend this class when you extend `ring_seminorm`. -/ -class ring_seminorm_class (F : Type*) (α : out_param $ Type*) [non_unital_non_assoc_ring α] - extends add_group_seminorm_class F α, submultiplicative_hom_class F α ℝ - -/-- `ring_norm_class F α` states that `F` is a type of norms on the ring `α`. - -You should extend this class when you extend `ring_norm`. -/ -class ring_norm_class (F : Type*) (α : out_param $ Type*) [non_unital_non_assoc_ring α] - extends ring_seminorm_class F α, add_group_norm_class F α - -/-- `mul_ring_seminorm_class F α` states that `F` is a type of multiplicative seminorms on the ring -`α`. - -You should extend this class when you extend `mul_ring_seminorm`. -/ -class mul_ring_seminorm_class (F : Type*) (α : out_param $ Type*) [non_assoc_ring α] - extends add_group_seminorm_class F α, monoid_with_zero_hom_class F α ℝ - -/-- `mul_ring_norm_class F α` states that `F` is a type of multiplicative norms on the ring `α`. - -You should extend this class when you extend `mul_ring_norm`. -/ -class mul_ring_norm_class (F : Type*) (α : out_param $ Type*) [non_assoc_ring α] - extends mul_ring_seminorm_class F α, add_group_norm_class F α - -@[priority 100] -- See note [lower instance priority] -instance mul_ring_seminorm_class.to_ring_seminorm_class [non_assoc_ring R] - [mul_ring_seminorm_class F R] : ring_seminorm_class F R := -{ map_mul_le_mul := λ f a b, (map_mul _ _ _).le, - ..‹mul_ring_seminorm_class F R› } - -@[priority 100] -- See note [lower instance priority] -instance mul_ring_norm_class.to_ring_norm_class [non_assoc_ring R] [mul_ring_norm_class F R] : - ring_norm_class F R := -{ ..‹mul_ring_norm_class F R›, ..mul_ring_seminorm_class.to_ring_seminorm_class } - namespace ring_seminorm section non_unital_ring variables [non_unital_ring R] -instance ring_seminorm_class : ring_seminorm_class (ring_seminorm R) R := +instance ring_seminorm_class : ring_seminorm_class (ring_seminorm R) R ℝ := { coe := λ f, f.to_fun, coe_injective' := λ f g h, by cases f; cases g; congr', map_zero := λ f, f.map_zero', @@ -177,7 +146,7 @@ namespace ring_norm variable [non_unital_ring R] -instance ring_norm_class : ring_norm_class (ring_norm R) R := +instance ring_norm_class : ring_norm_class (ring_norm R) R ℝ := { coe := λ f, f.to_fun, coe_injective' := λ f g h, by cases f; cases g; congr', map_zero := λ f, f.map_zero', @@ -210,7 +179,7 @@ end ring_norm namespace mul_ring_seminorm variables [non_assoc_ring R] -instance mul_ring_seminorm_class : mul_ring_seminorm_class (mul_ring_seminorm R) R := +instance mul_ring_seminorm_class : mul_ring_seminorm_class (mul_ring_seminorm R) R ℝ := { coe := λ f, f.to_fun, coe_injective' := λ f g h, by cases f; cases g; congr', map_zero := λ f, f.map_zero', @@ -250,7 +219,7 @@ end mul_ring_seminorm namespace mul_ring_norm variable [non_assoc_ring R] -instance mul_ring_norm_class : mul_ring_norm_class (mul_ring_norm R) R := +instance mul_ring_norm_class : mul_ring_norm_class (mul_ring_norm R) R ℝ := { coe := λ f, f.to_fun, coe_injective' := λ f g h, by cases f; cases g; congr', map_zero := λ f, f.map_zero', diff --git a/src/analysis/seminorm.lean b/src/analysis/seminorm.lean index 1fce620c351c3..e373ef838d830 100644 --- a/src/analysis/seminorm.lean +++ b/src/analysis/seminorm.lean @@ -52,7 +52,7 @@ attribute [nolint doc_blame] seminorm.to_add_group_seminorm You should extend this class when you extend `seminorm`. -/ class seminorm_class (F : Type*) (𝕜 E : out_param $ Type*) [semi_normed_ring 𝕜] [add_group E] - [has_smul 𝕜 E] extends add_group_seminorm_class F E := + [has_smul 𝕜 E] extends add_group_seminorm_class F E ℝ := (map_smul_eq_mul (f : F) (a : 𝕜) (x : E) : f (a • x) = ‖a‖ * f x) export seminorm_class (map_smul_eq_mul) From c8a43b4599e2a51b64db7b73f0527f19c3e5a1d0 Mon Sep 17 00:00:00 2001 From: Christopher Hoskin Date: Mon, 9 Jan 2023 19:56:48 +0000 Subject: [PATCH 0190/1291] feat(order/hom/complete_lattice): inf as an Inf_hom and sup as a Sup_hom (#18023) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Introduces a new definition `inf_Inf_hom` on a complete lattice which is the map `(a, b) ↦ a ⊓ b` as an `Inf_hom`. This is required for #17426. For completeness we also include the equivalent `sup` definition. Co-authored-by: Christopher Hoskin --- src/order/hom/complete_lattice.lean | 15 +++++++++++++++ 1 file changed, 15 insertions(+) diff --git a/src/order/hom/complete_lattice.lean b/src/order/hom/complete_lattice.lean index 7d3b73ead69ac..e37431e776b62 100644 --- a/src/order/hom/complete_lattice.lean +++ b/src/order/hom/complete_lattice.lean @@ -626,3 +626,18 @@ See also `complete_lattice_hom.set_preimage`. -/ right_inv := λ s, by simp only [← image_comp, equiv.self_comp_symm, id.def, image_id'], map_rel_iff' := λ s t, ⟨λ h, by simpa using @monotone_image _ _ e.symm _ _ h, λ h, monotone_image h⟩ } + +section + +variables (α) [complete_lattice α] + +/-- The map `(a, b) ↦ a ⊓ b` as an `Inf_hom`. -/ +@[simps] def inf_Inf_hom : Inf_hom (α × α) α := +{ to_fun := λ x, x.1 ⊓ x.2, + map_Inf' := λ s, by simp_rw [prod.fst_Inf, prod.snd_Inf, Inf_image, infi_inf_eq], } + +/-- The map `(a, b) ↦ a ⊔ b` as a `Sup_hom`. -/ +@[simps] def sup_Sup_hom : Sup_hom (α × α) α := +{ to_fun := λ x, x.1 ⊔ x.2, map_Sup' := (inf_Inf_hom αᵒᵈ).map_Inf' } + +end From 2f588be38bb5bec02f218ba14f82fc82eb663f87 Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Mon, 9 Jan 2023 19:56:49 +0000 Subject: [PATCH 0191/1291] refactor(ring_theory/fractional_ideal): refactor coe_to_fractional_ideal lemmas (#18055) Rename `coe_to_fractional_ideal` lemmas to `coe_ideal'` lemmas and extract `coe_ideal_inj` as standalone `eq_iff` lemmas. --- src/ring_theory/class_group.lean | 21 ++--- .../dedekind_domain/factorization.lean | 2 +- src/ring_theory/dedekind_domain/ideal.lean | 28 +++---- src/ring_theory/fractional_ideal.lean | 82 +++++++++---------- 4 files changed, 61 insertions(+), 72 deletions(-) diff --git a/src/ring_theory/class_group.lean b/src/ring_theory/class_group.lean index 6e894f437902d..cc46021ece831 100644 --- a/src/ring_theory/class_group.lean +++ b/src/ring_theory/class_group.lean @@ -159,8 +159,7 @@ by rw [class_group.mk, monoid_hom.comp_apply, ← monoid_hom.comp_apply (units.m /-- Send a nonzero integral ideal to an invertible fractional ideal. -/ noncomputable def fractional_ideal.mk0 [is_dedekind_domain R] : (ideal R)⁰ →* (fractional_ideal R⁰ K)ˣ := -{ to_fun := λ I, units.mk0 I ((fractional_ideal.coe_to_fractional_ideal_ne_zero (le_refl R⁰)).mpr - (mem_non_zero_divisors_iff_ne_zero.mp I.2)), +{ to_fun := λ I, units.mk0 I (coe_ideal_ne_zero.mpr $ mem_non_zero_divisors_iff_ne_zero.mp I.2), map_one' := by simp, map_mul' := λ x y, by simp } @@ -170,8 +169,7 @@ rfl lemma fractional_ideal.canonical_equiv_mk0 [is_dedekind_domain R] (K' : Type*) [field K'] [algebra R K'] [is_fraction_ring R K'] (I : (ideal R)⁰) : - fractional_ideal.canonical_equiv R⁰ K K' (fractional_ideal.mk0 K I) = - fractional_ideal.mk0 K' I := + fractional_ideal.canonical_equiv R⁰ K K' (fractional_ideal.mk0 K I) = fractional_ideal.mk0 K' I := by simp only [fractional_ideal.coe_mk0, coe_coe, fractional_ideal.canonical_equiv_coe_ideal] @[simp] lemma fractional_ideal.map_canonical_equiv_mk0 [is_dedekind_domain R] @@ -181,8 +179,7 @@ by simp only [fractional_ideal.coe_mk0, coe_coe, fractional_ideal.canonical_equi units.ext (fractional_ideal.canonical_equiv_mk0 K K' I) /-- Send a nonzero ideal to the corresponding class in the class group. -/ -noncomputable def class_group.mk0 [is_dedekind_domain R] : - (ideal R)⁰ →* class_group R := +noncomputable def class_group.mk0 [is_dedekind_domain R] : (ideal R)⁰ →* class_group R := class_group.mk.comp (fractional_ideal.mk0 (fraction_ring R)) @[simp] lemma class_group.mk_mk0 [is_dedekind_domain R] (I : (ideal R)⁰): @@ -202,8 +199,7 @@ begin end lemma class_group.mk0_eq_mk0_iff_exists_fraction_ring [is_dedekind_domain R] {I J : (ideal R)⁰} : - class_group.mk0 I = class_group.mk0 J ↔ - ∃ (x ≠ (0 : K)), span_singleton R⁰ x * I = J := + class_group.mk0 I = class_group.mk0 J ↔ ∃ (x ≠ (0 : K)), span_singleton R⁰ x * I = J := begin refine (class_group.equiv K).injective.eq_iff.symm.trans _, simp only [class_group.equiv_mk0, quotient_group.mk'_eq_mk', mem_principal_ideals_iff, @@ -305,8 +301,7 @@ begin simp [hx'] end -lemma class_group.mk0_eq_one_iff [is_dedekind_domain R] - {I : ideal R} (hI : I ∈ (ideal R)⁰) : +lemma class_group.mk0_eq_one_iff [is_dedekind_domain R] {I : ideal R} (hI : I ∈ (ideal R)⁰) : class_group.mk0 ⟨I, hI⟩ = 1 ↔ I.is_principal := class_group.mk_eq_one_iff.trans (coe_submodule_is_principal R _) @@ -315,8 +310,7 @@ class_group.mk_eq_one_iff.trans (coe_submodule_is_principal R _) See `class_group.fintype_of_admissible` for a finiteness proof that works for rings of integers of global fields. -/ -noncomputable instance [is_principal_ideal_ring R] : - fintype (class_group R) := +noncomputable instance [is_principal_ideal_ring R] : fintype (class_group R) := { elems := {1}, complete := begin @@ -326,8 +320,7 @@ noncomputable instance [is_principal_ideal_ring R] : end } /-- The class number of a principal ideal domain is `1`. -/ -lemma card_class_group_eq_one [is_principal_ideal_ring R] : - fintype.card (class_group R) = 1 := +lemma card_class_group_eq_one [is_principal_ideal_ring R] : fintype.card (class_group R) = 1 := begin rw fintype.card_eq_one_iff, use 1, diff --git a/src/ring_theory/dedekind_domain/factorization.lean b/src/ring_theory/dedekind_domain/factorization.lean index e2aa140725183..3527aeb7fcd0b 100644 --- a/src/ring_theory/dedekind_domain/factorization.lean +++ b/src/ring_theory/dedekind_domain/factorization.lean @@ -89,7 +89,7 @@ lemma finite_mul_support_coe {I : ideal R} (hI : I ≠ 0) : begin rw mul_support, simp_rw [ne.def, zpow_coe_nat, ← fractional_ideal.coe_ideal_pow, - fractional_ideal.coe_ideal_eq_one_iff], + fractional_ideal.coe_ideal_eq_one], exact finite_mul_support hI, end diff --git a/src/ring_theory/dedekind_domain/ideal.lean b/src/ring_theory/dedekind_domain/ideal.lean index 157f23c61278e..a54b392a6011e 100644 --- a/src/ring_theory/dedekind_domain/ideal.lean +++ b/src/ring_theory/dedekind_domain/ideal.lean @@ -266,8 +266,7 @@ begin refine is_noetherian_ring_iff.mpr ⟨λ (I : ideal A), _⟩, by_cases hI : I = ⊥, { rw hI, apply submodule.fg_bot }, - have hI : (I : fractional_ideal A⁰ (fraction_ring A)) ≠ 0 := - (coe_to_fractional_ideal_ne_zero (le_refl (non_zero_divisors A))).mpr hI, + have hI : (I : fractional_ideal A⁰ (fraction_ring A)) ≠ 0 := coe_ideal_ne_zero.mpr hI, exact I.fg_of_is_unit (is_fraction_ring.injective A (fraction_ring A)) (h.is_unit hI) end @@ -292,11 +291,9 @@ begin rintros P P_ne hP, refine ideal.is_maximal_def.mpr ⟨hP.ne_top, λ M hM, _⟩, -- We may assume `P` and `M` (as fractional ideals) are nonzero. - have P'_ne : (P : fractional_ideal A⁰ (fraction_ring A)) ≠ 0 := - (coe_to_fractional_ideal_ne_zero (le_refl (non_zero_divisors A))).mpr P_ne, + have P'_ne : (P : fractional_ideal A⁰ (fraction_ring A)) ≠ 0 := coe_ideal_ne_zero.mpr P_ne, have M'_ne : (M : fractional_ideal A⁰ (fraction_ring A)) ≠ 0 := - (coe_to_fractional_ideal_ne_zero (le_refl (non_zero_divisors A))).mpr - (lt_of_le_of_lt bot_le hM).ne', + coe_ideal_ne_zero.mpr (lt_of_le_of_lt bot_le hM).ne', -- In particular, we'll show `M⁻¹ * P ≤ P` suffices : (M⁻¹ * P : fractional_ideal A⁰ (fraction_ring A)) ≤ P, @@ -390,7 +387,7 @@ begin have hM0 := (M.bot_lt_of_maximal hNF).ne', obtain ⟨x, hxM, hx1⟩ := this hM, refine ⟨x, inv_anti_mono _ _ ((coe_ideal_le_coe_ideal _).mpr hIM) hxM, hx1⟩; - apply coe_ideal_ne_zero; assumption }, + rw coe_ideal_ne_zero; assumption }, -- Let `a` be a nonzero element of `M` and `J` the ideal generated by `a`. intros M hM, @@ -414,7 +411,7 @@ begin (λ h, hbJ $ h.symm ▸ J.zero_mem), -- Then `b a⁻¹ : K` is in `M⁻¹` but not in `1`. refine ⟨algebra_map A K b * (algebra_map A K a)⁻¹, (mem_inv_iff _).mpr _, _⟩, - { exact (coe_to_fractional_ideal_ne_zero le_rfl).mpr hM0.ne' }, + { exact coe_ideal_ne_zero.mpr hM0.ne' }, { rintro y₀ hy₀, obtain ⟨y, h_Iy, rfl⟩ := (mem_coe_ideal _).mp hy₀, rw [mul_comm, ← mul_assoc, ← ring_hom.map_mul], @@ -436,7 +433,7 @@ end lemma one_mem_inv_coe_ideal {I : ideal A} (hI : I ≠ ⊥) : (1 : K) ∈ (I : fractional_ideal A⁰ K)⁻¹ := begin - rw mem_inv_iff (coe_ideal_ne_zero hI), + rw mem_inv_iff (coe_ideal_ne_zero.mpr hI), intros y hy, rw one_mul, exact coe_ideal_le_one hy, @@ -496,7 +493,7 @@ begin have x_mul_mem : ∀ b ∈ (I⁻¹ : fractional_ideal A⁰ K), x * b ∈ (I⁻¹ : fractional_ideal A⁰ K), { intros b hb, rw mem_inv_iff at ⊢ hx, - swap, { exact coe_ideal_ne_zero hI0 }, + swap, { exact coe_ideal_ne_zero.mpr hI0 }, swap, { exact hJ0 }, simp only [mul_assoc, mul_comm b] at ⊢ hx, intros y hy, @@ -534,7 +531,7 @@ begin span_singleton_mul_span_singleton, inv_mul_cancel, span_singleton_one], { exact mt ((injective_iff_map_eq_zero (algebra_map A K)).mp (is_fraction_ring.injective A K) _) ha }, - { exact coe_ideal_ne_zero_iff.mp (right_ne_zero_of_mul hne) } + { exact coe_ideal_ne_zero.mp (right_ne_zero_of_mul hne) } end lemma mul_right_le_iff [is_dedekind_domain A] {J : fractional_ideal A⁰ K} @@ -601,7 +598,7 @@ noncomputable instance fractional_ideal.semifield : div := (/), div_eq_mul_inv := fractional_ideal.div_eq_mul_inv, mul_inv_cancel := λ I, fractional_ideal.mul_inv_cancel, - .. fractional_ideal.comm_semiring, ..(coe_to_fractional_ideal_injective le_rfl).nontrivial } + .. fractional_ideal.comm_semiring, .. coe_ideal_injective.nontrivial } /-- Fractional ideals have cancellative multiplication in a Dedekind domain. @@ -632,15 +629,14 @@ lemma ideal.dvd_iff_le {I J : ideal A} : (I ∣ J) ↔ J ≤ I := by_cases hI : I = ⊥, { have hJ : J = ⊥, { rwa [hI, ← eq_bot_iff] at h }, rw [hI, hJ] }, - have hI' : (I : fractional_ideal A⁰ (fraction_ring A)) ≠ 0 := - (coe_to_fractional_ideal_ne_zero (le_refl (non_zero_divisors A))).mpr hI, + have hI' : (I : fractional_ideal A⁰ (fraction_ring A)) ≠ 0 := coe_ideal_ne_zero.mpr hI, have : (I : fractional_ideal A⁰ (fraction_ring A))⁻¹ * J ≤ 1 := le_trans (mul_left_mono (↑I)⁻¹ ((coe_ideal_le_coe_ideal _).mpr h)) (le_of_eq (inv_mul_cancel hI')), obtain ⟨H, hH⟩ := le_one_iff_exists_coe_ideal.mp this, use H, - refine coe_to_fractional_ideal_injective (le_refl (non_zero_divisors A)) - (show (J : fractional_ideal A⁰ (fraction_ring A)) = _, from _), + refine coe_ideal_injective + (show (J : fractional_ideal A⁰ (fraction_ring A)) = ↑(I * H), from _), rw [coe_ideal_mul, hH, ← mul_assoc, mul_inv_cancel hI', one_mul] end⟩ diff --git a/src/ring_theory/fractional_ideal.lean b/src/ring_theory/fractional_ideal.lean index 42cb8547169ec..c17ff5cdfe2bc 100644 --- a/src/ring_theory/fractional_ideal.lean +++ b/src/ring_theory/fractional_ideal.lean @@ -108,8 +108,7 @@ variables [algebra R P] [loc : is_localization S P] This coercion is typically called `coe_to_submodule` in lemma names (or `coe` when the coercion is clear from the context), not to be confused with `is_localization.coe_submodule : ideal R → submodule R P` -(which we use to define `coe : ideal R → fractional_ideal S P`, -referred to as `coe_ideal` in theorem names). +(which we use to define `coe : ideal R → fractional_ideal S P`). -/ instance : has_coe (fractional_ideal S P) (submodule R P) := ⟨λ I, I.val⟩ @@ -156,8 +155,10 @@ lemma coe_to_submodule_injective : function.injective (coe : fractional_ideal S P → submodule R P) := subtype.coe_injective -lemma is_fractional_of_le_one (I : submodule R P) (h : I ≤ 1) : - is_fractional S I := +lemma coe_to_submodule_inj {I J : fractional_ideal S P} : (I : submodule R P) = J ↔ I = J := +coe_to_submodule_injective.eq_iff + +lemma is_fractional_of_le_one (I : submodule R P) (h : I ≤ 1) : is_fractional S I := begin use [1, S.one_mem], intros b hb, @@ -166,8 +167,8 @@ begin exact set.mem_range_self b', end -lemma is_fractional_of_le {I : submodule R P} {J : fractional_ideal S P} - (hIJ : I ≤ J) : is_fractional S I := +lemma is_fractional_of_le {I : submodule R P} {J : fractional_ideal S P} (hIJ : I ≤ J) : + is_fractional S I := begin obtain ⟨a, a_mem, ha⟩ := J.is_fractional, use [a, a_mem], @@ -184,9 +185,9 @@ also called `coe_to_submodule` in theorem names. This map is available as a ring hom, called `fractional_ideal.coe_ideal_hom`. -/ -- Is a `coe_t` rather than `coe` to speed up failing inference, see library note [use has_coe_t] -instance coe_to_fractional_ideal : has_coe_t (ideal R) (fractional_ideal S P) := -⟨λ I, ⟨coe_submodule P I, is_fractional_of_le_one _ - (by simpa using coe_submodule_mono P (le_top : I ≤ ⊤))⟩⟩ +instance : has_coe_t (ideal R) (fractional_ideal S P) := +⟨λ I, ⟨coe_submodule P I, + is_fractional_of_le_one _ $ by simpa using coe_submodule_mono P (le_top : I ≤ ⊤)⟩⟩ @[simp, norm_cast] lemma coe_coe_ideal (I : ideal R) : ((I : fractional_ideal S P) : submodule R P) = coe_submodule P I := rfl @@ -222,8 +223,7 @@ variables {S} @[simp, norm_cast] lemma coe_zero : ↑(0 : fractional_ideal S P) = (⊥ : submodule R P) := submodule.ext $ λ _, mem_zero_iff S -@[simp, norm_cast] lemma coe_to_fractional_ideal_bot : ((⊥ : ideal R) : fractional_ideal S P) = 0 := -rfl +@[simp, norm_cast] lemma coe_ideal_bot : ((⊥ : ideal R) : fractional_ideal S P) = 0 := rfl variables (P) @@ -235,22 +235,21 @@ include loc variables {P} -lemma coe_to_fractional_ideal_injective (h : S ≤ non_zero_divisors R) : +lemma coe_ideal_injective' (h : S ≤ non_zero_divisors R) : function.injective (coe : ideal R → fractional_ideal S P) := -λ I J heq, have - ∀ (x : R), algebra_map R P x ∈ (I : fractional_ideal S P) ↔ - algebra_map R P x ∈ (J : fractional_ideal S P) := -λ x, heq ▸ iff.rfl, -ideal.ext (by simpa only [mem_coe_ideal, exists_prop, exists_mem_to_map_eq P h] using this) +λ _ _ h', ((coe_ideal_le_coe_ideal' S h).mp h'.le).antisymm ((coe_ideal_le_coe_ideal' S h).mp h'.ge) -lemma coe_to_fractional_ideal_eq_zero {I : ideal R} (hS : S ≤ non_zero_divisors R) : +lemma coe_ideal_inj' (h : S ≤ non_zero_divisors R) {I J : ideal R} : + (I : fractional_ideal S P) = J ↔ I = J := +(coe_ideal_injective' h).eq_iff + +@[simp] lemma coe_ideal_eq_zero' {I : ideal R} (h : S ≤ non_zero_divisors R) : (I : fractional_ideal S P) = 0 ↔ I = (⊥ : ideal R) := -⟨λ h, coe_to_fractional_ideal_injective hS h, - λ h, by rw [h, coe_to_fractional_ideal_bot]⟩ +coe_ideal_inj' h -lemma coe_to_fractional_ideal_ne_zero {I : ideal R} (hS : S ≤ non_zero_divisors R) : +lemma coe_ideal_ne_zero' {I : ideal R} (h : S ≤ non_zero_divisors R) : (I : fractional_ideal S P) ≠ 0 ↔ I ≠ (⊥ : ideal R) := -not_iff_not.mpr (coe_to_fractional_ideal_eq_zero hS) +not_iff_not.mpr $ coe_ideal_eq_zero' h omit loc @@ -270,8 +269,7 @@ instance : has_one (fractional_ideal S P) := variables (S) -@[simp, norm_cast] lemma coe_ideal_top : ((⊤ : ideal R) : fractional_ideal S P) = 1 := -rfl +@[simp, norm_cast] lemma coe_ideal_top : ((⊤ : ideal R) : fractional_ideal S P) = 1 := rfl lemma mem_one_iff {x : P} : x ∈ (1 : fractional_ideal S P) ↔ ∃ x' : R, algebra_map R P x' = x := iff.intro (λ ⟨x', _, h⟩, ⟨x', h⟩) (λ ⟨x', h⟩, ⟨x', ⟨⟩, h⟩) @@ -538,7 +536,7 @@ def coe_ideal_hom : ideal R →+* fractional_ideal S P := map_add' := coe_ideal_sup, map_mul' := coe_ideal_mul, map_one' := by rw [ideal.one_eq_top, coe_ideal_top], - map_zero' := coe_to_fractional_ideal_bot } + map_zero' := coe_ideal_bot } lemma coe_ideal_pow (I : ideal R) (n : ℕ) : (↑(I^n) : fractional_ideal S P) = I^n := (coe_ideal_hom S P).map_pow _ n @@ -547,8 +545,7 @@ open_locale big_operators lemma coe_ideal_finprod [is_localization S P] {α : Sort*} {f : α → ideal R} (hS : S ≤ non_zero_divisors R) : ((∏ᶠ a : α, f a : ideal R) : fractional_ideal S P) = ∏ᶠ a : α, (f a : fractional_ideal S P) := -monoid_hom.map_finprod_of_injective (coe_ideal_hom S P).to_monoid_hom - (coe_to_fractional_ideal_injective hS) f +monoid_hom.map_finprod_of_injective (coe_ideal_hom S P).to_monoid_hom (coe_ideal_injective' hS) f end order @@ -815,19 +812,23 @@ end ⟨imp_of_not_imp_not _ _ (map_ne_zero _), λ hI, hI.symm ▸ map_zero h⟩ lemma coe_ideal_injective : function.injective (coe : ideal R → fractional_ideal R⁰ K) := -injective_of_le_imp_le _ (λ _ _, (coe_ideal_le_coe_ideal _).mp) +coe_ideal_injective' le_rfl + +lemma coe_ideal_inj {I J : ideal R} : + (I : fractional_ideal R⁰ K) = (J : fractional_ideal R⁰ K) ↔ I = J := +coe_ideal_inj' le_rfl -@[simp] lemma coe_ideal_eq_zero_iff {I : ideal R} : (I : fractional_ideal R⁰ K) = 0 ↔ I = ⊥ := -coe_ideal_injective.eq_iff +@[simp] lemma coe_ideal_eq_zero {I : ideal R} : (I : fractional_ideal R⁰ K) = 0 ↔ I = ⊥ := +coe_ideal_eq_zero' le_rfl -lemma coe_ideal_ne_zero_iff {I : ideal R} : (I : fractional_ideal R⁰ K) ≠ 0 ↔ I ≠ ⊥ := -not_iff_not.mpr coe_ideal_eq_zero_iff +lemma coe_ideal_ne_zero {I : ideal R} : (I : fractional_ideal R⁰ K) ≠ 0 ↔ I ≠ ⊥ := +coe_ideal_ne_zero' le_rfl -lemma coe_ideal_ne_zero {I : ideal R} (hI : I ≠ ⊥) : (I : fractional_ideal R⁰ K) ≠ 0 := -coe_ideal_ne_zero_iff.mpr hI +@[simp] lemma coe_ideal_eq_one {I : ideal R} : (I : fractional_ideal R⁰ K) = 1 ↔ I = 1 := +by simpa only [ideal.one_eq_top] using coe_ideal_inj -@[simp] lemma coe_ideal_eq_one_iff {I : ideal R} : (I : fractional_ideal R⁰ K) = 1 ↔ I = 1 := -by simpa only [ideal.one_eq_top] using coe_ideal_injective.eq_iff +lemma coe_ideal_ne_one {I : ideal R} : (I : fractional_ideal R⁰ K) ≠ 1 ↔ I ≠ 1 := +not_iff_not.mpr coe_ideal_eq_one end is_fraction_ring @@ -1049,7 +1050,7 @@ variables {R₁} @[simp] lemma span_finset_eq_zero {ι : Type*} {s : finset ι} {f : ι → K} : span_finset R₁ s f = 0 ↔ ∀ j ∈ s, f j = 0 := -by simp only [← coe_to_submodule_injective.eq_iff, span_finset_coe, coe_zero, submodule.span_eq_bot, +by simp only [← coe_to_submodule_inj, span_finset_coe, coe_zero, submodule.span_eq_bot, set.mem_image, finset.mem_coe, forall_exists_index, and_imp, forall_apply_eq_imp_iff₂] lemma span_finset_ne_zero {ι : Type*} {s : finset ι} {f : ι → K} : @@ -1197,14 +1198,13 @@ lemma mk'_mul_coe_ideal_eq_coe_ideal {I J : ideal R₁} {x y : R₁} (hy : y ∈ span_singleton R₁⁰ (is_localization.mk' K x ⟨y, hy⟩) * I = (J : fractional_ideal R₁⁰ K) ↔ ideal.span {x} * I = ideal.span {y} * J := begin - have inj : function.injective (coe : ideal R₁ → fractional_ideal R₁⁰ K) := coe_ideal_injective, have : span_singleton R₁⁰ (is_localization.mk' _ (1 : R₁) ⟨y, hy⟩) * span_singleton R₁⁰ (algebra_map R₁ K y) = 1, { rw [span_singleton_mul_span_singleton, mul_comm, ← is_localization.mk'_eq_mul_mk'_one, is_localization.mk'_self, span_singleton_one] }, let y' : (fractional_ideal R₁⁰ K)ˣ := units.mk_of_mul_eq_one _ _ this, have coe_y' : ↑y' = span_singleton R₁⁰ (is_localization.mk' K (1 : R₁) ⟨y, hy⟩) := rfl, - refine iff.trans _ (y'.mul_right_inj.trans inj.eq_iff), + refine iff.trans _ (y'.mul_right_inj.trans coe_ideal_inj), rw [coe_y', coe_ideal_mul, coe_ideal_span_singleton, coe_ideal_mul, coe_ideal_span_singleton, ←mul_assoc, span_singleton_mul_span_singleton, ←mul_assoc, span_singleton_mul_span_singleton, mul_comm (mk' _ _ _), ← is_localization.mk'_eq_mul_mk'_one, @@ -1323,7 +1323,7 @@ lemma is_noetherian_iff {I : fractional_ideal R₁⁰ K} : is_noetherian R₁ I ↔ ∀ J ≤ I, (J : submodule R₁ K).fg := is_noetherian_submodule.trans ⟨λ h J hJ, h _ hJ, λ h J hJ, h ⟨J, is_fractional_of_le hJ⟩ hJ⟩ -lemma is_noetherian_coe_to_fractional_ideal [_root_.is_noetherian_ring R₁] (I : ideal R₁) : +lemma is_noetherian_coe_ideal [_root_.is_noetherian_ring R₁] (I : ideal R₁) : is_noetherian R₁ (I : fractional_ideal R₁⁰ K) := begin rw is_noetherian_iff, @@ -1364,7 +1364,7 @@ theorem is_noetherian [_root_.is_noetherian_ring R₁] (I : fractional_ideal R begin obtain ⟨d, J, h_nzd, rfl⟩ := exists_eq_span_singleton_mul I, apply is_noetherian_span_singleton_inv_to_map_mul, - apply is_noetherian_coe_to_fractional_ideal, + apply is_noetherian_coe_ideal end section adjoin From 2f7b36ab1dffd22fcc18e8ad84340e7314abc819 Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Mon, 9 Jan 2023 19:56:50 +0000 Subject: [PATCH 0192/1291] feat(linear_algebra/span): add submodule.mem_span lemmas (#18100) --- src/linear_algebra/span.lean | 3 +++ src/ring_theory/ideal/basic.lean | 2 +- src/ring_theory/principal_ideal_domain.lean | 8 ++++---- 3 files changed, 8 insertions(+), 5 deletions(-) diff --git a/src/linear_algebra/span.lean b/src/linear_algebra/span.lean index 2735fed2ba80e..004f7e0b618bf 100644 --- a/src/linear_algebra/span.lean +++ b/src/linear_algebra/span.lean @@ -426,6 +426,9 @@ begin simp only [eq_comm, add_comm, exists_and_distrib_left] end +lemma mem_span_pair {x y z : M} : z ∈ span R ({x, y} : set M) ↔ ∃ a b : R, a • x + b • y = z := +by simp_rw [mem_span_insert, mem_span_singleton, exists_prop, exists_exists_eq_and, eq_comm] + lemma span_insert (x) (s : set M) : span R (insert x s) = span R ({x} : set M) ⊔ span R s := by rw [insert_eq, span_union] diff --git a/src/ring_theory/ideal/basic.lean b/src/ring_theory/ideal/basic.lean index 9769b03b72100..fd1ffa938c20c 100644 --- a/src/ring_theory/ideal/basic.lean +++ b/src/ring_theory/ideal/basic.lean @@ -272,7 +272,7 @@ by simp only [span_insert, sup_comm] theorem mem_span_pair {x y z : α} : z ∈ span ({x, y} : set α) ↔ ∃ a b, a * x + b * y = z := -by simp [mem_span_insert, mem_span_singleton', @eq_comm _ _ z] +submodule.mem_span_pair @[simp] lemma span_pair_add_mul_left {R : Type u} [comm_ring R] {x y : R} (z : R) : (span {x + y * z, y} : ideal R) = span {x, y} := diff --git a/src/ring_theory/principal_ideal_domain.lean b/src/ring_theory/principal_ideal_domain.lean index 4906cdcf429a9..ccf8a7be9eacb 100644 --- a/src/ring_theory/principal_ideal_domain.lean +++ b/src/ring_theory/principal_ideal_domain.lean @@ -308,19 +308,19 @@ begin exact (span_singleton_mul_right_unit D.is_unit _) }, apply associated_of_dvd_dvd, { rw dvd_gcd_iff, - split; rw [←ideal.mem_span_singleton, ←hd, mem_span_pair], + split; rw [←ideal.mem_span_singleton, ←hd, ideal.mem_span_pair], { use [1, 0], rw [one_mul, zero_mul, add_zero] }, { use [0, 1], rw [one_mul, zero_mul, zero_add] } }, { obtain ⟨r, s, rfl⟩ : ∃ r s, r * x + s * y = d, - { rw [←mem_span_pair, hd, ideal.mem_span_singleton] }, + { rw [←ideal.mem_span_pair, hd, ideal.mem_span_singleton] }, apply dvd_add; apply dvd_mul_of_dvd_right, exacts [gcd_dvd_left x y, gcd_dvd_right x y] }, end theorem gcd_dvd_iff_exists (a b : R) {z} : gcd a b ∣ z ↔ ∃ x y, z = a * x + b * y := -by simp_rw [mul_comm a, mul_comm b, @eq_comm _ z, ←mem_span_pair, ←span_gcd, +by simp_rw [mul_comm a, mul_comm b, @eq_comm _ z, ←ideal.mem_span_pair, ←span_gcd, ideal.mem_span_singleton] /-- **Bézout's lemma** -/ @@ -328,7 +328,7 @@ theorem exists_gcd_eq_mul_add_mul (a b : R) : ∃ x y, gcd a b = a * x + b * y : by rw [←gcd_dvd_iff_exists] theorem gcd_is_unit_iff (x y : R) : is_unit (gcd x y) ↔ is_coprime x y := -by rw [is_coprime, ←mem_span_pair, ←span_gcd, ←span_singleton_eq_top, eq_top_iff_one] +by rw [is_coprime, ←ideal.mem_span_pair, ←span_gcd, ←span_singleton_eq_top, eq_top_iff_one] -- this should be proved for UFDs surely? theorem is_coprime_of_dvd (x y : R) From 09d74faae35ef6197bb2dba4d0ea7939b9440ead Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Tue, 10 Jan 2023 06:03:25 +0000 Subject: [PATCH 0193/1291] feat(ring_theory/class_group): add class_group.mk_eq_mk (#18034) --- src/ring_theory/class_group.lean | 50 ++++++++++++++++++++-- src/ring_theory/dedekind_domain/ideal.lean | 6 +-- 2 files changed, 50 insertions(+), 6 deletions(-) diff --git a/src/ring_theory/class_group.lean b/src/ring_theory/class_group.lean index cc46021ece831..5e913eba6f342 100644 --- a/src/ring_theory/class_group.lean +++ b/src/ring_theory/class_group.lean @@ -74,19 +74,19 @@ begin { exact ⟨x, hx⟩ }, { refine ⟨units.mk0 x _, hx⟩, rintro rfl, - simpa [I.ne_zero.symm] using hx }, - + simpa [I.ne_zero.symm] using hx } end instance principal_ideals.normal : (to_principal_ideal R K).range.normal := subgroup.normal_of_comm _ + end variables (R) [is_domain R] /-- The ideal class group of `R` is the group of invertible fractional ideals modulo the principal ideals. -/ -@[derive(comm_group)] +@[derive comm_group] def class_group := (fractional_ideal R⁰ (fraction_ring R))ˣ ⧸ (to_principal_ideal R (fraction_ring R)).range @@ -99,6 +99,50 @@ noncomputable def class_group.mk : (fractional_ideal R⁰ K)ˣ →* class_group (quotient_group.mk' (to_principal_ideal R (fraction_ring R)).range).comp (units.map (fractional_ideal.canonical_equiv R⁰ K (fraction_ring R))) +lemma class_group.mk_eq_mk {I J : (fractional_ideal R⁰ $ fraction_ring R)ˣ} : + class_group.mk I = class_group.mk J + ↔ ∃ x : (fraction_ring R)ˣ, I * to_principal_ideal R (fraction_ring R) x = J := +by { erw [quotient_group.mk'_eq_mk', canonical_equiv_self, units.map_id, set.exists_range_iff], + refl } + +lemma class_group.mk_eq_mk_of_coe_ideal + {I J : (fractional_ideal R⁰ $ fraction_ring R)ˣ} {I' J' : ideal R} + (hI : (I : fractional_ideal R⁰ $ fraction_ring R) = I') + (hJ : (J : fractional_ideal R⁰ $ fraction_ring R) = J') : + class_group.mk I = class_group.mk J + ↔ ∃ x y : R, x ≠ 0 ∧ y ≠ 0 ∧ ideal.span {x} * I' = ideal.span {y} * J' := +begin + rw [class_group.mk_eq_mk], + split, + { rintro ⟨x, rfl⟩, + rw [units.coe_mul, hI, coe_to_principal_ideal, mul_comm, + span_singleton_mul_coe_ideal_eq_coe_ideal] at hJ, + exact ⟨_, _, sec_fst_ne_zero le_rfl x.ne_zero, sec_snd_ne_zero le_rfl ↑x, hJ⟩ }, + { rintro ⟨x, y, hx, hy, h⟩, + split, rw [mul_comm, ← units.eq_iff, units.coe_mul, coe_to_principal_ideal], + convert (mk'_mul_coe_ideal_eq_coe_ideal (fraction_ring R) $ + mem_non_zero_divisors_of_ne_zero hy).2 h, + apply (ne.is_unit _).unit_spec, + rwa [ne, mk'_eq_zero_iff_eq_zero] } +end + +lemma class_group.mk_eq_one_of_coe_ideal {I : (fractional_ideal R⁰ $ fraction_ring R)ˣ} + {I' : ideal R} (hI : (I : fractional_ideal R⁰ $ fraction_ring R) = I') : + class_group.mk I = 1 ↔ ∃ x : R, x ≠ 0 ∧ I' = ideal.span {x} := +begin + rw [← map_one class_group.mk, class_group.mk_eq_mk_of_coe_ideal hI (_ : _ = ↑⊤)], + any_goals { refl }, + split, + { rintro ⟨x, y, hx, hy, h⟩, + rw [ideal.mul_top] at h, + rcases ideal.mem_span_singleton_mul.mp ((ideal.span_singleton_le_iff_mem _).mp h.ge) + with ⟨i, hi, rfl⟩, + rw [← ideal.span_singleton_mul_span_singleton, ideal.span_singleton_mul_right_inj hx] at h, + exact ⟨i, right_ne_zero_of_mul hy, h⟩ }, + { rintro ⟨x, hx, rfl⟩, + exact ⟨1, x, one_ne_zero, hx, by rw [ideal.span_singleton_one, ideal.top_mul, ideal.mul_top]⟩ } +end + variables (K) /-- Induction principle for the class group: to show something holds for all `x : class_group R`, diff --git a/src/ring_theory/dedekind_domain/ideal.lean b/src/ring_theory/dedekind_domain/ideal.lean index a54b392a6011e..9adff7d3d09cf 100644 --- a/src/ring_theory/dedekind_domain/ideal.lean +++ b/src/ring_theory/dedekind_domain/ideal.lean @@ -138,7 +138,7 @@ lemma span_singleton_div_self {x : K} (hx : x ≠ 0) : by rw [span_singleton_div_span_singleton, div_self hx, span_singleton_one] lemma coe_ideal_span_singleton_div_self {x : R₁} (hx : x ≠ 0) : - ((ideal.span {x} : ideal R₁) : fractional_ideal R₁⁰ K) / (ideal.span {x} : ideal R₁) = 1 := + (ideal.span ({x} : set R₁) : fractional_ideal R₁⁰ K) / ideal.span ({x} : set R₁) = 1 := by rw [coe_ideal_span_singleton, span_singleton_div_self K $ (map_ne_zero_iff _ $ no_zero_smul_divisors.algebra_map_injective R₁ K).mpr hx] @@ -147,7 +147,7 @@ lemma span_singleton_mul_inv {x : K} (hx : x ≠ 0) : by rw [span_singleton_inv, span_singleton_mul_span_singleton, mul_inv_cancel hx, span_singleton_one] lemma coe_ideal_span_singleton_mul_inv {x : R₁} (hx : x ≠ 0) : - ((ideal.span {x} : ideal R₁) : fractional_ideal R₁⁰ K) * (ideal.span {x} : ideal R₁)⁻¹ = 1 := + (ideal.span ({x} : set R₁) : fractional_ideal R₁⁰ K) * (ideal.span ({x} : set R₁))⁻¹ = 1 := by rw [coe_ideal_span_singleton, span_singleton_mul_inv K $ (map_ne_zero_iff _ $ no_zero_smul_divisors.algebra_map_injective R₁ K).mpr hx] @@ -156,7 +156,7 @@ lemma span_singleton_inv_mul {x : K} (hx : x ≠ 0) : by rw [mul_comm, span_singleton_mul_inv K hx] lemma coe_ideal_span_singleton_inv_mul {x : R₁} (hx : x ≠ 0) : - ((ideal.span {x} : ideal R₁) : fractional_ideal R₁⁰ K)⁻¹ * (ideal.span {x} : ideal R₁) = 1 := + (ideal.span ({x} : set R₁) : fractional_ideal R₁⁰ K)⁻¹ * ideal.span ({x} : set R₁) = 1 := by rw [mul_comm, coe_ideal_span_singleton_mul_inv K hx] lemma mul_generator_self_inv {R₁ : Type*} [comm_ring R₁] [algebra R₁ K] [is_localization R₁⁰ K] From 7b78d1776212a91ecc94cf601f83bdcc46b04213 Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Tue, 10 Jan 2023 07:38:25 +0000 Subject: [PATCH 0194/1291] chore(group_theory/presented_group): generalize universe (#18115) closes #18114 --- src/group_theory/presented_group.lean | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/src/group_theory/presented_group.lean b/src/group_theory/presented_group.lean index e1c6eb448ef22..fcd3117a13e30 100644 --- a/src/group_theory/presented_group.lean +++ b/src/group_theory/presented_group.lean @@ -25,11 +25,11 @@ given by generators `x : α` and relations `r ∈ rels`. generators, relations, group presentations -/ -variables {α : Type} +variables {α : Type*} /-- Given a set of relations, rels, over a type `α`, presented_group constructs the group with generators `x : α` and relations `rels` as a quotient of free_group `α`.-/ -def presented_group (rels : set (free_group α)) : Type := +def presented_group (rels : set (free_group α)) := free_group α ⧸ subgroup.normal_closure rels namespace presented_group @@ -50,7 +50,7 @@ the images of `f` satisfy all the given relations, then `f` extends uniquely to from `presented_group rels` to `G`. -/ -variables {G : Type} [group G] {f : α → G} {rels : set (free_group α)} +variables {G : Type*} [group G] {f : α → G} {rels : set (free_group α)} local notation `F` := free_group.lift f From e017e666c7182d23a6385459c3b36ae453a5be34 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 10 Jan 2023 09:28:58 +0000 Subject: [PATCH 0195/1291] refactor(algebra/order/to_interval_mod): state more results in terms of interval sets (#18102) Since the declarations contain `Ioc` or `Ico` in their name, I would argue the statements are clearer if expressed with the corresponding sets rather than with inequalities. While this makes a few proofs longer in terms of characters, they're shorter in terms of terms. --- src/algebra/order/to_interval_mod.lean | 64 +++++++++++++------------- 1 file changed, 33 insertions(+), 31 deletions(-) diff --git a/src/algebra/order/to_interval_mod.lean b/src/algebra/order/to_interval_mod.lean index bc775f63cb434..3a7b1cb9762a4 100644 --- a/src/algebra/order/to_interval_mod.lean +++ b/src/algebra/order/to_interval_mod.lean @@ -136,32 +136,30 @@ by rw [←neg_sub, to_Ico_mod_sub_to_Ico_div_zsmul] by rw [←neg_sub, to_Ioc_mod_sub_to_Ioc_div_zsmul] lemma to_Ico_mod_eq_iff {a b x y : α} (hb : 0 < b) : - to_Ico_mod a hb x = y ↔ a ≤ y ∧ y < a + b ∧ ∃ z : ℤ, y - x = z • b := + to_Ico_mod a hb x = y ↔ y ∈ set.Ico a (a + b) ∧ ∃ z : ℤ, y - x = z • b := begin - refine ⟨λ h, ⟨h ▸ left_le_to_Ico_mod a hb x, - h ▸ to_Ico_mod_lt_right a hb x, + refine ⟨λ h, ⟨h ▸ to_Ico_mod_mem_Ico a hb x, to_Ico_div a hb x, h ▸ to_Ico_mod_sub_self a hb x⟩, λ h, _⟩, - rcases h with ⟨ha, hab, z, hz⟩, + rcases h with ⟨hy, z, hz⟩, rw sub_eq_iff_eq_add' at hz, subst hz, - rw eq_to_Ico_div_of_add_zsmul_mem_Ico hb (set.mem_Ico.2 ⟨ha, hab⟩), + rw eq_to_Ico_div_of_add_zsmul_mem_Ico hb hy, refl end lemma to_Ioc_mod_eq_iff {a b x y : α} (hb : 0 < b) : - to_Ioc_mod a hb x = y ↔ a < y ∧ y ≤ a + b ∧ ∃ z : ℤ, y - x = z • b := + to_Ioc_mod a hb x = y ↔ y ∈ set.Ioc a (a + b) ∧ ∃ z : ℤ, y - x = z • b := begin - refine ⟨λ h, ⟨h ▸ left_lt_to_Ioc_mod a hb x, - h ▸ to_Ioc_mod_le_right a hb x, + refine ⟨λ h, ⟨h ▸ to_Ioc_mod_mem_Ioc a hb x, to_Ioc_div a hb x, h ▸ to_Ioc_mod_sub_self a hb x⟩, λ h, _⟩, - rcases h with ⟨ha, hab, z, hz⟩, + rcases h with ⟨hy, z, hz⟩, rw sub_eq_iff_eq_add' at hz, subst hz, - rw eq_to_Ioc_div_of_add_zsmul_mem_Ioc hb (set.mem_Ioc.2 ⟨ha, hab⟩), + rw eq_to_Ioc_div_of_add_zsmul_mem_Ioc hb hy, refl end @@ -179,15 +177,15 @@ end @[simp] lemma to_Ico_mod_apply_left (a : α) {b : α} (hb : 0 < b) : to_Ico_mod a hb a = a := begin - rw to_Ico_mod_eq_iff hb, - refine ⟨le_refl _, lt_add_of_pos_right _ hb, 0, _⟩, + rw [to_Ico_mod_eq_iff hb, set.left_mem_Ico], + refine ⟨lt_add_of_pos_right _ hb, 0, _⟩, simp end @[simp] lemma to_Ioc_mod_apply_left (a : α) {b : α} (hb : 0 < b) : to_Ioc_mod a hb a = a + b := begin - rw to_Ioc_mod_eq_iff hb, - refine ⟨lt_add_of_pos_right _ hb, le_refl _, 1, _⟩, + rw [to_Ioc_mod_eq_iff hb, set.right_mem_Ioc], + refine ⟨lt_add_of_pos_right _ hb, 1, _⟩, simp end @@ -207,16 +205,16 @@ end lemma to_Ico_mod_apply_right (a : α) {b : α} (hb : 0 < b) : to_Ico_mod a hb (a + b) = a := begin - rw to_Ico_mod_eq_iff hb, - refine ⟨le_refl _, lt_add_of_pos_right _ hb, -1, _⟩, + rw [to_Ico_mod_eq_iff hb, set.left_mem_Ico], + refine ⟨lt_add_of_pos_right _ hb, -1, _⟩, simp end lemma to_Ioc_mod_apply_right (a : α) {b : α} (hb : 0 < b) : to_Ioc_mod a hb (a + b) = a + b := begin - rw to_Ioc_mod_eq_iff hb, - refine ⟨lt_add_of_pos_right _ hb, le_refl _, 0, _⟩, + rw [to_Ioc_mod_eq_iff hb, set.right_mem_Ioc], + refine ⟨lt_add_of_pos_right _ hb, 0, _⟩, simp end @@ -476,13 +474,15 @@ lemma tfae_mem_Ioo_mod : to_Ico_mod a hb x ≠ a] := begin tfae_have : 1 → 2, - { exact λ ⟨i, hi⟩, ((to_Ico_mod_eq_iff hb).2 ⟨hi.1.le, hi.2, i, add_sub_cancel' x _⟩).trans - ((to_Ioc_mod_eq_iff hb).2 ⟨hi.1, hi.2.le, i, add_sub_cancel' x _⟩).symm }, + { exact λ ⟨i, hi⟩, + ((to_Ico_mod_eq_iff hb).2 ⟨set.Ioo_subset_Ico_self hi, i, add_sub_cancel' x _⟩).trans + ((to_Ioc_mod_eq_iff hb).2 ⟨set.Ioo_subset_Ioc_self hi, i, add_sub_cancel' x _⟩).symm }, tfae_have : 2 → 3, { intro h, rw [h, ne, add_right_eq_self], exact hb.ne' }, tfae_have : 3 → 4, - { refine mt (λ h, _), rw [h, eq_comm, to_Ioc_mod_eq_iff], - refine ⟨lt_add_of_pos_right a hb, le_rfl, to_Ico_div a hb x + 1, _⟩, + { refine mt (λ h, _), + rw [h, eq_comm, to_Ioc_mod_eq_iff, set.right_mem_Ioc], + refine ⟨lt_add_of_pos_right a hb, to_Ico_div a hb x + 1, _⟩, conv_lhs { rw [← h, to_Ico_mod, add_assoc, ← add_one_zsmul, add_sub_cancel'] } }, tfae_have : 4 → 1, { have h' := to_Ico_mod_mem_Ico a hb x, exact λ h, ⟨_, h'.1.lt_of_ne' h, h'.2⟩ }, @@ -501,7 +501,7 @@ lemma mem_Ioo_mod_iff_to_Ioc_mod_ne_right : mem_Ioo_mod a b x ↔ to_Ioc_mod a h begin rw [mem_Ioo_mod_iff_to_Ico_mod_eq_to_Ioc_mod, to_Ico_mod_eq_iff hb], obtain ⟨h₁, h₂⟩ := to_Ioc_mod_mem_Ioc a hb x, - exact ⟨λ h, h.2.1.ne, λ h, ⟨h₁.le, h₂.lt_of_ne h, _, add_sub_cancel' x _⟩⟩, + exact ⟨λ h, h.1.2.ne, λ h, ⟨⟨h₁.le, h₂.lt_of_ne h⟩, _, add_sub_cancel' x _⟩⟩, end lemma mem_Ioo_mod_iff_to_Ico_div_eq_to_Ioc_div : @@ -520,8 +520,10 @@ lemma mem_Ioo_mod_iff_sub_ne_zsmul : mem_Ioo_mod a b x ↔ ∀ z : ℤ, a - x begin rw [mem_Ioo_mod_iff_to_Ico_mod_ne_left hb, ← not_iff_not], push_neg, split; intro h, - { rw ← h, exact ⟨_, add_sub_cancel' x _⟩ }, - { exact (to_Ico_mod_eq_iff hb).2 ⟨le_rfl, lt_add_of_pos_right a hb, h⟩ }, + { rw ← h, + exact ⟨_, add_sub_cancel' x _⟩ }, + { rw [to_Ico_mod_eq_iff, set.left_mem_Ico], + exact ⟨lt_add_of_pos_right a hb, h⟩, }, end lemma mem_Ioo_mod_iff_eq_mod_zmultiples : @@ -541,17 +543,17 @@ end end Ico_Ioc -lemma to_Ico_mod_eq_self {a b x : α} (hb : 0 < b) : to_Ico_mod a hb x = x ↔ a ≤ x ∧ x < a + b := +lemma to_Ico_mod_eq_self {a b x : α} (hb : 0 < b) : to_Ico_mod a hb x = x ↔ x ∈ set.Ico a (a + b) := begin - rw to_Ico_mod_eq_iff, - refine ⟨λ h, ⟨h.1, h.2.1⟩, λ h, ⟨h.1, h.2, 0, _⟩⟩, + rw [to_Ico_mod_eq_iff, and_iff_left], + refine ⟨0, _⟩, simp end -lemma to_Ioc_mod_eq_self {a b x : α} (hb : 0 < b) : to_Ioc_mod a hb x = x ↔ a < x ∧ x ≤ a + b := +lemma to_Ioc_mod_eq_self {a b x : α} (hb : 0 < b) : to_Ioc_mod a hb x = x ↔ x ∈ set.Ioc a (a + b) := begin - rw to_Ioc_mod_eq_iff, - refine ⟨λ h, ⟨h.1, h.2.1⟩, λ h, ⟨h.1, h.2, 0, _⟩⟩, + rw [to_Ioc_mod_eq_iff, and_iff_left], + refine ⟨0, _⟩, simp end From 6f413f3f7330b94c92a5a27488fdc74e6d483a78 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 10 Jan 2023 11:05:26 +0000 Subject: [PATCH 0196/1291] chore(algebra/order/archimedean): add `sub` variations of `add` lemmas (#18103) These spellings give the intuition of "reduce"ing to a range rather than "lifting" to a range. --- src/algebra/order/archimedean.lean | 19 +++++++++++++++---- 1 file changed, 15 insertions(+), 4 deletions(-) diff --git a/src/algebra/order/archimedean.lean b/src/algebra/order/archimedean.lean index bf4cdcc97a2ec..0212343cf1717 100644 --- a/src/algebra/order/archimedean.lean +++ b/src/algebra/order/archimedean.lean @@ -68,19 +68,30 @@ lemma exists_unique_zsmul_near_of_pos' {a : α} (ha : 0 < a) (g : α) : by simpa only [sub_nonneg, add_zsmul, one_zsmul, sub_lt_iff_lt_add'] using exists_unique_zsmul_near_of_pos ha g +lemma exists_unique_sub_zsmul_mem_Ico {a : α} (ha : 0 < a) (b c : α) : + ∃! m : ℤ, b - m • a ∈ set.Ico c (c + a) := +by simpa only [mem_Ico, le_sub_iff_add_le, zero_add, add_comm c, sub_lt_iff_lt_add', add_assoc] + using exists_unique_zsmul_near_of_pos' ha (b - c) + lemma exists_unique_add_zsmul_mem_Ico {a : α} (ha : 0 < a) (b c : α) : ∃! m : ℤ, b + m • a ∈ set.Ico c (c + a) := (equiv.neg ℤ).bijective.exists_unique_iff.2 $ - by simpa only [equiv.neg_apply, mem_Ico, neg_zsmul, ← sub_eq_add_neg, le_sub_iff_add_le, zero_add, - add_comm c, sub_lt_iff_lt_add', add_assoc] using exists_unique_zsmul_near_of_pos' ha (b - c) + by simpa only [equiv.neg_apply, neg_zsmul, ← sub_eq_add_neg] + using exists_unique_sub_zsmul_mem_Ico ha b c lemma exists_unique_add_zsmul_mem_Ioc {a : α} (ha : 0 < a) (b c : α) : ∃! m : ℤ, b + m • a ∈ set.Ioc c (c + a) := (equiv.add_right (1 : ℤ)).bijective.exists_unique_iff.2 $ - by simpa only [add_zsmul, sub_lt_iff_lt_add', le_sub_iff_add_le', ← add_assoc, and.comm, mem_Ioc, - equiv.coe_add_right, one_zsmul, add_le_add_iff_right] + by simpa only [add_one_zsmul, sub_lt_iff_lt_add', le_sub_iff_add_le', ← add_assoc, and.comm, + mem_Ioc, equiv.coe_add_right, add_le_add_iff_right] using exists_unique_zsmul_near_of_pos ha (c - b) +lemma exists_unique_sub_zsmul_mem_Ioc {a : α} (ha : 0 < a) (b c : α) : + ∃! m : ℤ, b - m • a ∈ set.Ioc c (c + a) := +(equiv.neg ℤ).bijective.exists_unique_iff.2 $ + by simpa only [equiv.neg_apply, neg_zsmul, sub_neg_eq_add] + using exists_unique_add_zsmul_mem_Ioc ha b c + end linear_ordered_add_comm_group theorem exists_nat_gt [strict_ordered_semiring α] [archimedean α] (x : α) : ∃ n : ℕ, x < n := From b86832321b586c6ac23ef8cdef6a7a27e42b13bd Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 10 Jan 2023 13:41:40 +0000 Subject: [PATCH 0197/1291] feat(group_theory/perm/cycle/basic): Cycle on a set (#17899) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Define `equiv.perm.is_cycle_on`, a predicate for a permutation to be a cycle on a set. This has two discrepancies with the existing `equiv.perm.is_cycle_on` (apart from the obvious one that `is_cycle_on` is restricted to a set): * `is_cycle_on` forbids fixed points (on `s`) while `is_cycle` allows them. * `is_cycle` forbids the identity while `is_cycle_on` allows it. I think the first one isn't so bad given that `is_cycle` is meant to be used over the entire type (so you can't just rule out fixed points). The second one is more of a worry as I had to case-split on `s.subsingleton ∨ s.nontrivial` in several `is_cycle_on` lemmas to derive them from the corresponding `is_cycle` ones, while of course the `is_cycle` ones would also have held for a weaker version of `is_cycle` that allows the identity. --- src/data/set/function.lean | 6 + src/dynamics/fixed_points/basic.lean | 2 +- src/group_theory/perm/basic.lean | 31 +- src/group_theory/perm/cycle/basic.lean | 330 ++++++++++++++++------ src/group_theory/perm/cycle/concrete.lean | 12 +- src/group_theory/perm/support.lean | 4 + 6 files changed, 284 insertions(+), 101 deletions(-) diff --git a/src/data/set/function.lean b/src/data/set/function.lean index b51f2b6bb7f0e..ad181e2038170 100644 --- a/src/data/set/function.lean +++ b/src/data/set/function.lean @@ -1449,6 +1449,12 @@ lemma inv_on : inv_on e e.symm t s := lemma bij_on_image : bij_on e s (e '' s) := (e.injective.inj_on _).bij_on_image lemma bij_on_symm_image : bij_on e.symm (e '' s) s := e.bij_on_image.symm e.inv_on +variables {e} + +@[simp] lemma bij_on_symm : bij_on e.symm t s ↔ bij_on e s t := bij_on_comm e.symm.inv_on + +alias bij_on_symm ↔ _root_.set.bij_on.of_equiv_symm _root_.set.bij_on.equiv_symm + variables [decidable_eq α] {a b : α} lemma bij_on_swap (ha : a ∈ s) (hb : b ∈ s) : bij_on (swap a b) s s := diff --git a/src/dynamics/fixed_points/basic.lean b/src/dynamics/fixed_points/basic.lean index 2520af1f44024..d5eda405f92dc 100644 --- a/src/dynamics/fixed_points/basic.lean +++ b/src/dynamics/fixed_points/basic.lean @@ -87,7 +87,7 @@ h.to_left_inverse e.left_inverse_symm protected lemma perm_inv (h : is_fixed_pt e x) : is_fixed_pt ⇑(e⁻¹) x := h.equiv_symm protected lemma perm_pow (h : is_fixed_pt e x) (n : ℕ) : is_fixed_pt ⇑(e ^ n) x := -by { rw ←equiv.perm.iterate_eq_pow, exact h.iterate _ } +by { rw equiv.perm.coe_pow, exact h.iterate _ } protected lemma perm_zpow (h : is_fixed_pt e x) : ∀ n : ℤ, is_fixed_pt ⇑(e ^ n) x | (int.of_nat n) := h.perm_pow _ diff --git a/src/group_theory/perm/basic.lean b/src/group_theory/perm/basic.lean index 7ae9350c243e3..0f6b7a4f8c9ab 100644 --- a/src/group_theory/perm/basic.lean +++ b/src/group_theory/perm/basic.lean @@ -6,6 +6,7 @@ Authors: Leonardo de Moura, Mario Carneiro import algebra.group.pi import algebra.group.prod import algebra.hom.iterate +import logic.equiv.set /-! # The group of permutations (self-equivalences) of a type `α` @@ -67,9 +68,9 @@ lemma inv_def (f : perm α) : f⁻¹ = f.symm := rfl @[simp, norm_cast] lemma coe_one : ⇑(1 : perm α) = id := rfl @[simp, norm_cast] lemma coe_mul (f g : perm α) : ⇑(f * g) = f ∘ g := rfl - @[norm_cast] lemma coe_pow (f : perm α) (n : ℕ) : ⇑(f ^ n) = (f^[n]) := hom_coe_pow _ rfl (λ _ _, rfl) _ _ +@[simp] lemma iterate_eq_pow (f : perm α) (n : ℕ) : (f^[n]) = ⇑(f ^ n) := (coe_pow _ _).symm lemma eq_inv_iff_eq {f : perm α} {x y : α} : x = f⁻¹ y ↔ f x = y := f.eq_symm_apply @@ -79,9 +80,9 @@ lemma zpow_apply_comm {α : Type*} (σ : perm α) (m n : ℤ) {x : α} : (σ ^ m) ((σ ^ n) x) = (σ ^ n) ((σ ^ m) x) := by rw [←equiv.perm.mul_apply, ←equiv.perm.mul_apply, zpow_mul_comm] -@[simp] lemma iterate_eq_pow (f : perm α) : ∀ n, f^[n] = ⇑(f ^ n) -| 0 := rfl -| (n + 1) := by { rw [function.iterate_succ, pow_add, iterate_eq_pow], refl } +@[simp] lemma image_inv (f : perm α) (s : set α) : ⇑f⁻¹ '' s = f ⁻¹' s := f⁻¹.image_eq_preimage _ +@[simp] lemma preimage_inv (f : perm α) (s : set α) : ⇑f⁻¹ ⁻¹' s = f '' s := +(f.image_eq_preimage _).symm /-! Lemmas about mixing `perm` with `equiv`. Because we have multiple ways to express `equiv.refl`, `equiv.symm`, and `equiv.trans`, we want simp lemmas for every combination. @@ -494,3 +495,25 @@ lemma zpow_mul_right : ∀ n : ℤ, equiv.mul_right a ^ n = equiv.mul_right (a ^ end group end equiv + +open equiv function + +namespace set +variables {α : Type*} {f : perm α} {s t : set α} + +@[simp] lemma bij_on_perm_inv : bij_on ⇑f⁻¹ t s ↔ bij_on f s t := equiv.bij_on_symm + +alias bij_on_perm_inv ↔ bij_on.of_perm_inv bij_on.perm_inv + +lemma maps_to.perm_pow : maps_to f s s → ∀ n : ℕ, maps_to ⇑(f ^ n) s s := +by { simp_rw equiv.perm.coe_pow, exact maps_to.iterate } +lemma surj_on.perm_pow : surj_on f s s → ∀ n : ℕ, surj_on ⇑(f ^ n) s s := +by { simp_rw equiv.perm.coe_pow, exact surj_on.iterate } +lemma bij_on.perm_pow : bij_on f s s → ∀ n : ℕ, bij_on ⇑(f ^ n) s s := +by { simp_rw equiv.perm.coe_pow, exact bij_on.iterate } + +lemma bij_on.perm_zpow (hf : bij_on f s s) : ∀ n : ℤ, bij_on ⇑(f ^ n) s s +| (int.of_nat n) := hf.perm_pow _ +| (int.neg_succ_of_nat n) := by { rw zpow_neg_succ_of_nat, exact (hf.perm_pow _).perm_inv } + +end set diff --git a/src/group_theory/perm/cycle/basic.lean b/src/group_theory/perm/cycle/basic.lean index c0860c622aa6d..caa1530b9dcef 100644 --- a/src/group_theory/perm/cycle/basic.lean +++ b/src/group_theory/perm/cycle/basic.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2019 Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Chris Hughes +Authors: Chris Hughes, Yaël Dillies -/ import data.finset.noncomm_prod import data.fintype.perm @@ -11,13 +11,17 @@ import logic.equiv.fintype /-! # Cyclic permutations +This file develops the theory of cycles in permutations. + ## Main definitions In the following, `f : equiv.perm β`. * `equiv.perm.same_cycle`: `f.same_cycle x y` when `x` and `y` are in the same cycle of `f`. -* `equiv.perm.is_cycle`: `f.is_cycle` when two nonfixed points of `β` - are related by repeated application of `f`. +* `equiv.perm.is_cycle`: `f` is a cycle if any two nonfixed points of `f` are related by repeated + applications of `f`, and `f` is not the identity. +* `equiv.perm.is_cycle_on`: `f` is a cycle on a set `s` when any two points of `s` are related by + repeated applications of `f`. The following two definitions require that `β` is a `fintype`: @@ -36,6 +40,12 @@ The following two definitions require that `β` is a `fintype`: - `closure_prime_cycle_swap` : The symmetric group is generated by a prime cycle and a transposition +## Notes + +`equiv.perm.is_cycle` and `equiv.perm.is_cycle_on` are different in three ways: +* `is_cycle` is about the entire type while `is_cycle_on` is restricted to a set. +* `is_cycle` forbids the identity while `is_cycle_on` allows it (if `s` is a subsingleton). +* `is_cycle_on` forbids fixed points on `s` (if `s` is nontrivial), while `is_cycle` allows them. -/ open equiv function finset @@ -100,22 +110,26 @@ by rw [←same_cycle_apply_left, apply_inv_self] @[simp] lemma same_cycle_inv_apply_right : same_cycle f x (f⁻¹ y) ↔ same_cycle f x y := by rw [←same_cycle_apply_right, apply_inv_self] -@[simp] lemma same_cycle_zpow_left_iff {n : ℤ} : same_cycle f ((f ^ n) x) y ↔ same_cycle f x y := +@[simp] lemma same_cycle_zpow_left {n : ℤ} : same_cycle f ((f ^ n) x) y ↔ same_cycle f x y := (equiv.add_right (n : ℤ)).exists_congr_left.trans $ by simp [same_cycle, zpow_add] -@[simp] lemma same_cycle_zpow_right_iff {n : ℤ} : same_cycle f x ((f ^ n) y) ↔ same_cycle f x y := -by rw [same_cycle_comm, same_cycle_zpow_left_iff, same_cycle_comm] +@[simp] lemma same_cycle_zpow_right {n : ℤ} : same_cycle f x ((f ^ n) y) ↔ same_cycle f x y := +by rw [same_cycle_comm, same_cycle_zpow_left, same_cycle_comm] -@[simp] lemma same_cycle_pow_left_iff {n : ℕ} : same_cycle f ((f ^ n) x) y ↔ same_cycle f x y := -by rw [←zpow_coe_nat, same_cycle_zpow_left_iff] +@[simp] lemma same_cycle_pow_left {n : ℕ} : same_cycle f ((f ^ n) x) y ↔ same_cycle f x y := +by rw [←zpow_coe_nat, same_cycle_zpow_left] -@[simp] lemma same_cycle_pow_right_iff {n : ℕ} : same_cycle f x ((f ^ n) y) ↔ same_cycle f x y := -by rw [←zpow_coe_nat, same_cycle_zpow_right_iff] +@[simp] lemma same_cycle_pow_right {n : ℕ} : same_cycle f x ((f ^ n) y) ↔ same_cycle f x y := +by rw [←zpow_coe_nat, same_cycle_zpow_right] -alias same_cycle_pow_left_iff ↔ _ same_cycle.pow_left -alias same_cycle_pow_right_iff ↔ _ same_cycle.pow_right -alias same_cycle_zpow_left_iff ↔ _ same_cycle.zpow_left -alias same_cycle_zpow_right_iff ↔ _ same_cycle.zpow_right +alias same_cycle_apply_left ↔ same_cycle.of_apply_left same_cycle.apply_left +alias same_cycle_apply_right ↔ same_cycle.of_apply_right same_cycle.apply_right +alias same_cycle_inv_apply_left ↔ same_cycle.of_inv_apply_left same_cycle.inv_apply_left +alias same_cycle_inv_apply_right ↔ same_cycle.of_inv_apply_right same_cycle.inv_apply_right +alias same_cycle_pow_left ↔ same_cycle.of_pow_left same_cycle.pow_left +alias same_cycle_pow_right ↔ same_cycle.of_pow_right same_cycle.pow_right +alias same_cycle_zpow_left ↔ same_cycle.of_zpow_left same_cycle.zpow_left +alias same_cycle_zpow_right ↔ same_cycle.of_zpow_right same_cycle.zpow_right lemma same_cycle.of_pow {n : ℕ} : same_cycle (f ^ n) x y → same_cycle f x y := λ ⟨m, h⟩, ⟨n * m, by simp [zpow_mul, h]⟩ @@ -136,7 +150,7 @@ exists_congr $ λ n, by rw [←extend_domain_zpow, extend_domain_apply_image, su alias same_cycle_extend_domain ↔ _ same_cycle.extend_domain -lemma same_cycle.nat' [finite α] : same_cycle f x y → ∃ i < order_of f, (f ^ i) x = y := +lemma same_cycle.exists_pow_eq' [finite α] : same_cycle f x y → ∃ i < order_of f, (f ^ i) x = y := begin classical, rintro ⟨k, rfl⟩, @@ -149,11 +163,11 @@ begin exact int.mod_lt_of_pos _ h₀, end -lemma same_cycle.nat'' [finite α] (h : same_cycle f x y) : +lemma same_cycle.exists_pow_eq'' [finite α] (h : same_cycle f x y) : ∃ (i : ℕ) (hpos : 0 < i) (h : i ≤ order_of f), (f ^ i) x = y := begin classical, - obtain ⟨_|i, hi, rfl⟩ := h.nat', + obtain ⟨_ | i, hi, rfl⟩ := h.exists_pow_eq', { refine ⟨order_of f, order_of_pos f, le_rfl, _⟩, rw [pow_order_of_eq_one, pow_zero] }, { exact ⟨i.succ, i.zero_lt_succ, hi.le, rfl⟩ } @@ -605,6 +619,177 @@ end end is_cycle +/-! ### `is_cycle_on` -/ + +section is_cycle_on +variables {f g : perm α} {s t : set α} {a b x y : α} + +/-- A permutation is a cycle on `s` when any two points of `s` are related by repeated application +of the permutation. Note that this means the identity is a cycle of subsingleton sets. -/ +def is_cycle_on (f : perm α) (s : set α) : Prop := +set.bij_on f s s ∧ ∀ ⦃x⦄, x ∈ s → ∀ ⦃y⦄, y ∈ s → f.same_cycle x y + +@[simp] lemma is_cycle_on_empty : f.is_cycle_on ∅ := by simp [is_cycle_on] + +@[simp] lemma is_cycle_on_one : (1 : perm α).is_cycle_on s ↔ s.subsingleton := +by simp [is_cycle_on, set.bij_on_id, set.subsingleton] + +alias is_cycle_on_one ↔ is_cycle_on.subsingleton _root_.set.subsingleton.is_cycle_on_one + +@[simp] lemma is_cycle_on_singleton : f.is_cycle_on {a} ↔ f a = a := +by simp [is_cycle_on, same_cycle.rfl] + +lemma is_cycle_on_of_subsingleton [subsingleton α] (f : perm α) (s : set α) : f.is_cycle_on s := +⟨s.bij_on_of_subsingleton _, λ x _ y _, (subsingleton.elim x y).same_cycle _⟩ + +@[simp] lemma is_cycle_on_inv : f⁻¹.is_cycle_on s ↔ f.is_cycle_on s := +by simp [is_cycle_on, set.bij_on_perm_inv] + +alias is_cycle_on_inv ↔ is_cycle_on.of_inv is_cycle_on.inv + +lemma is_cycle_on.conj (h : f.is_cycle_on s) : (g * f * g⁻¹).is_cycle_on ((g : perm α) '' s) := +⟨(g.bij_on_image.comp h.1).comp g.bij_on_symm_image, + λ x hx y hy, by { rw ←preimage_inv at hx hy, convert (h.2 hx hy).conj; rw apply_inv_self }⟩ + +lemma is_cycle_on_swap [decidable_eq α] (hab : a ≠ b) : (swap a b).is_cycle_on {a, b} := +⟨bij_on_swap (by simp) (by simp), λ x hx y hy, begin + rw [set.mem_insert_iff, set.mem_singleton_iff] at hx hy, + obtain rfl | rfl := hx; obtain rfl | rfl := hy, + { exact ⟨0, by rw [zpow_zero, coe_one, id.def]⟩ }, + { exact ⟨1, by rw [zpow_one, swap_apply_left]⟩ }, + { exact ⟨1, by rw [zpow_one, swap_apply_right]⟩ }, + { exact ⟨0, by rw [zpow_zero, coe_one, id.def]⟩ } +end⟩ + +protected lemma is_cycle_on.apply_ne (hf : f.is_cycle_on s) (hs : s.nontrivial) (ha : a ∈ s) : + f a ≠ a := +begin + obtain ⟨b, hb, hba⟩ := hs.exists_ne a, + obtain ⟨n, rfl⟩ := hf.2 ha hb, + exact λ h, hba (is_fixed_pt.perm_zpow h n), +end + +protected lemma is_cycle.is_cycle_on (hf : f.is_cycle) : f.is_cycle_on {x | f x ≠ x} := +⟨f.bij_on $ λ x, f.apply_eq_iff_eq.not, λ a ha b, hf.same_cycle ha⟩ + +/-- This lemma demonstrates the relation between `equiv.perm.is_cycle` and `equiv.perm.is_cycle_on` +in non-degenerate cases. -/ +lemma is_cycle_iff_exists_is_cycle_on : + f.is_cycle ↔ ∃ s : set α, s.nontrivial ∧ f.is_cycle_on s ∧ ∀ ⦃x⦄, ¬ is_fixed_pt f x → x ∈ s := +begin + refine ⟨λ hf, ⟨{x | f x ≠ x}, _, hf.is_cycle_on, λ _, id⟩, _⟩, + { obtain ⟨a, ha⟩ := hf, + exact ⟨f a, f.injective.ne ha.1, a, ha.1, ha.1⟩ }, + { rintro ⟨s, hs, hf, hsf⟩, + obtain ⟨a, ha⟩ := hs.nonempty, + exact ⟨a, hf.apply_ne hs ha, λ b hb, hf.2 ha $ hsf hb⟩ } +end + +lemma is_cycle_on.apply_mem_iff (hf : f.is_cycle_on s) : f x ∈ s ↔ x ∈ s := +⟨λ hx, by { convert hf.1.perm_inv.1 hx, rw inv_apply_self }, λ hx, hf.1.maps_to hx⟩ + +/-- Note that the identity satisfies `is_cycle_on` for any subsingleton set, but not `is_cycle`. -/ +lemma is_cycle_on.is_cycle_subtype_perm (hf : f.is_cycle_on s) (hs : s.nontrivial) : + (f.subtype_perm $ λ _, hf.apply_mem_iff.symm : perm s).is_cycle := +begin + obtain ⟨a, ha⟩ := hs.nonempty, + exact ⟨⟨a, ha⟩, ne_of_apply_ne (coe : s → α) (hf.apply_ne hs ha), + λ b hb, (hf.2 (⟨a, ha⟩ : s).prop b.prop).subtype_perm⟩, +end + +/-- Note that the identity is a cycle on any subsingleton set, but not a cycle. -/ +protected lemma is_cycle_on.subtype_perm (hf : f.is_cycle_on s) : + (f.subtype_perm $ λ _, hf.apply_mem_iff.symm : perm s).is_cycle_on set.univ := +begin + obtain hs | hs := s.subsingleton_or_nontrivial, + { haveI := hs.coe_sort, + exact is_cycle_on_of_subsingleton _ _ }, + convert (hf.is_cycle_subtype_perm hs).is_cycle_on, + rw [eq_comm, set.eq_univ_iff_forall], + exact λ x, ne_of_apply_ne (coe : s → α) (hf.apply_ne hs x.prop), +end + +-- TODO: Theory of order of an element under an action +lemma is_cycle_on.pow_apply_eq {s : finset α} (hf : f.is_cycle_on s) (ha : a ∈ s) {n : ℕ} : + (f ^ n) a = a ↔ s.card ∣ n := +begin + obtain rfl | hs := finset.eq_singleton_or_nontrivial ha, + { rw [coe_singleton, is_cycle_on_singleton] at hf, + simpa using is_fixed_pt.iterate hf n }, + classical, + have h : ∀ x ∈ s.attach, ¬ f ↑x = ↑x := λ x hx, hf.apply_ne hs x.prop, + have := (hf.is_cycle_subtype_perm hs).order_of, + simp only [filter_true_of_mem h, support_subtype_perm, card_attach] at this, + rw [←this, order_of_dvd_iff_pow_eq_one, (hf.is_cycle_subtype_perm hs).pow_eq_one_iff' + (ne_of_apply_ne (coe : s → α) $ hf.apply_ne hs (⟨a, ha⟩ : s).prop)], + simp only [subtype.coe_mk, subtype_perm_pow, subtype_perm_apply], +end + +lemma is_cycle_on.zpow_apply_eq {s : finset α} (hf : f.is_cycle_on s) (ha : a ∈ s) : + ∀ {n : ℤ}, (f ^ n) a = a ↔ (s.card : ℤ) ∣ n +| (int.of_nat n) := (hf.pow_apply_eq ha).trans int.coe_nat_dvd.symm +| (int.neg_succ_of_nat n) := by { rw [zpow_neg_succ_of_nat, ←inv_pow], + exact (hf.inv.pow_apply_eq ha).trans ((dvd_neg _ _).trans int.coe_nat_dvd).symm } + +lemma is_cycle_on.pow_apply_eq_pow_apply {s : finset α} (hf : f.is_cycle_on s) (ha : a ∈ s) + {m n : ℕ} : (f ^ m) a = (f ^ n) a ↔ m ≡ n [MOD s.card] := +begin + rw [nat.modeq_iff_dvd, ←hf.zpow_apply_eq ha], + simp [sub_eq_neg_add, zpow_add, eq_inv_iff_eq, eq_comm], +end + +lemma is_cycle_on.zpow_apply_eq_zpow_apply {s : finset α} (hf : f.is_cycle_on s) (ha : a ∈ s) + {m n : ℤ} : (f ^ m) a = (f ^ n) a ↔ m ≡ n [ZMOD s.card] := +begin + rw [int.modeq_iff_dvd, ←hf.zpow_apply_eq ha], + simp [sub_eq_neg_add, zpow_add, eq_inv_iff_eq, eq_comm], +end + +lemma is_cycle_on.pow_card_apply {s : finset α} (hf : f.is_cycle_on s) (ha : a ∈ s) : + (f ^ s.card) a = a := +(hf.pow_apply_eq ha).2 dvd_rfl + +lemma is_cycle_on.exists_pow_eq {s : finset α} (hf : f.is_cycle_on s) (ha : a ∈ s) (hb : b ∈ s) : + ∃ n < s.card, (f ^ n) a = b := +begin + classical, + obtain ⟨n, rfl⟩ := hf.2 ha hb, + obtain ⟨k, hk⟩ := (int.mod_modeq n s.card).symm.dvd, + refine ⟨n.nat_mod s.card, int.nat_mod_lt (nonempty.card_pos ⟨a, ha⟩).ne', _⟩, + rw [←zpow_coe_nat, int.nat_mod, int.to_nat_of_nonneg (int.mod_nonneg _ $ nat.cast_ne_zero.2 + (nonempty.card_pos ⟨a, ha⟩).ne'), sub_eq_iff_eq_add'.1 hk, zpow_add, zpow_mul], + simp only [zpow_coe_nat, coe_mul, embedding_like.apply_eq_iff_eq], + exact is_fixed_pt.perm_zpow (hf.pow_card_apply ha) _, +end + +lemma is_cycle_on.exists_pow_eq' (hs : s.finite) (hf : f.is_cycle_on s) (ha : a ∈ s) (hb : b ∈ s) : + ∃ n : ℕ, (f ^ n) a = b := +by { lift s to finset α using id hs, obtain ⟨n, -, hn⟩ := hf.exists_pow_eq ha hb, exact ⟨n, hn⟩ } + +lemma is_cycle_on.range_pow (hs : s.finite) (h : f.is_cycle_on s) (ha : a ∈ s) : + set.range (λ n, (f ^ n) a : ℕ → α) = s := +set.subset.antisymm (set.range_subset_iff.2 $ λ n, h.1.maps_to.perm_pow _ ha) $ + λ x, h.exists_pow_eq' hs ha + +lemma is_cycle_on.range_zpow (h : f.is_cycle_on s) (ha : a ∈ s) : + set.range (λ n, (f ^ n) a : ℤ → α) = s := +set.subset.antisymm (set.range_subset_iff.2 $ λ n, (h.1.perm_zpow _).maps_to ha) $ h.2 ha + +lemma is_cycle_on.of_pow {n : ℕ} (hf : (f ^ n).is_cycle_on s) (h : set.bij_on f s s) : + f.is_cycle_on s := +⟨h, λ x hx y hy, (hf.2 hx hy).of_pow⟩ + +lemma is_cycle_on.of_zpow {n : ℤ} (hf : (f ^ n).is_cycle_on s) (h : set.bij_on f s s) : + f.is_cycle_on s := +⟨h, λ x hx y hy, (hf.2 hx hy).of_zpow⟩ + +lemma is_cycle_on.extend_domain {p : β → Prop} [decidable_pred p] (f : α ≃ subtype p) + (h : g.is_cycle_on s) : + (g.extend_domain f).is_cycle_on (coe ∘ f '' s) := +⟨h.1.extend_domain, by { rintro _ ⟨a, ha, rfl⟩ _ ⟨b, hb, rfl⟩, exact (h.2 ha hb).extend_domain }⟩ + +end is_cycle_on + /-! ### `cycle_of` -/ @@ -802,14 +987,45 @@ begin ←pow_eq_mod_order_of] } end -/-- x is in the support of f iff cycle_of f x is a cycle.-/ -lemma is_cycle_cycle_of_iff (f : perm α) : is_cycle (cycle_of f x) ↔ (f x ≠ x) := +/-- `x` is in the support of `f` iff `equiv.perm.cycle_of f x` is a cycle. -/ +lemma is_cycle_cycle_of_iff (f : perm α) : is_cycle (cycle_of f x) ↔ f x ≠ x := begin - split, - { intro hx, rw ne.def, rw ← cycle_of_eq_one_iff f, - exact equiv.perm.is_cycle.ne_one hx, }, - { intro hx, - apply equiv.perm.is_cycle_cycle_of, exact hx } + refine ⟨λ hx, _, f.is_cycle_cycle_of⟩, + rw [ne.def, ←cycle_of_eq_one_iff f], + exact hx.ne_one, +end + +lemma is_cycle_on_support_cycle_of (f : perm α) (x : α) : f.is_cycle_on (f.cycle_of x).support := +⟨f.bij_on $ by simp [mem_support_cycle_of_iff], λ a ha b hb, + by { rw [mem_coe, mem_support_cycle_of_iff] at ha hb, exact ha.1.symm.trans hb.1 }⟩ + +lemma same_cycle.exists_pow_eq_of_mem_support (h : same_cycle f x y) (hx : x ∈ f.support) : + ∃ (i : ℕ) (hi' : i < (f.cycle_of x).support.card), (f ^ i) x = y := +begin + rw mem_support at hx, + refine (f.is_cycle_on_support_cycle_of _).exists_pow_eq _ _; + rwa mem_support_cycle_of_iff' hx, +end + +lemma same_cycle.exists_pow_eq (f : perm α) (h : same_cycle f x y) : + ∃ (i : ℕ) (hi : 0 < i) (hi' : i ≤ (f.cycle_of x).support.card + 1), (f ^ i) x = y := +begin + by_cases hx : x ∈ f.support, + { obtain ⟨k, hk, hk'⟩ := h.exists_pow_eq_of_mem_support hx, + cases k, + { refine ⟨(f.cycle_of x).support.card, _, self_le_add_right _ _, _⟩, + { refine zero_lt_one.trans (one_lt_card_support_of_ne_one _), + simpa using hx }, + { simp only [perm.coe_one, id.def, pow_zero] at hk', + subst hk', + rw [←(is_cycle_cycle_of _ $ mem_support.1 hx).order_of, + ←cycle_of_pow_apply_self, pow_order_of_eq_one, one_apply] } }, + { exact ⟨k + 1, by simp, nat.le_succ_of_le hk.le, hk'⟩ } }, + { refine ⟨1, zero_lt_one, by simp, _⟩, + obtain ⟨k, rfl⟩ := h, + rw [not_mem_support] at hx, + rw [pow_apply_eq_self_of_apply_eq_self hx, + zpow_apply_eq_self_of_apply_eq_self hx] } end end cycle_of @@ -1171,72 +1387,6 @@ begin { exact λ H, hd.disjoint_cycle_factors_finset.le_bot (mem_inter_of_mem H hf) } } } end -lemma same_cycle.nat_of_mem_support [fintype α] (f : perm α) {x y : α} (h : same_cycle f x y) - (hx : x ∈ f.support) : - ∃ (i : ℕ) (hi' : i < (f.cycle_of x).support.card), (f ^ i) x = y := -begin - revert f, - intro f, - apply cycle_induction_on _ f, - { simp }, - { intros g hg H hx, - rw mem_support at hx, - rw [hg.cycle_of_eq hx, ←hg.order_of], - exact H.nat' }, - { rintros g h hd hg IH IH' ⟨m, rfl⟩ hx, - cases (disjoint_iff_eq_or_eq.mp hd) x with hgx hhx, - { have hpow : ∀ (k : ℤ), ((g * h) ^ k) x = (h ^ k) x, - { intro k, - suffices : (g ^ k) x = x, - { simpa [hd.commute.eq, hd.commute.symm.mul_zpow] }, - rw zpow_apply_eq_self_of_apply_eq_self, - simpa using hgx }, - obtain ⟨k, hk, hk'⟩ := IH' _ _, - { refine ⟨k, _, _⟩, - { rw [←cycle_of_eq_one_iff] at hgx, - rwa [hd.cycle_of_mul_distrib, hgx, one_mul] }, - { simpa [←zpow_coe_nat, hpow] using hk' } }, - { use m, - simp [hpow] }, - { rw [mem_support, hd.commute.eq] at hx, - simpa [hgx] using hx } }, - { have hpow : ∀ (k : ℤ), ((g * h) ^ k) x = (g ^ k) x, - { intro k, - suffices : (h ^ k) x = x, - { simpa [hd.commute.mul_zpow] }, - rw zpow_apply_eq_self_of_apply_eq_self, - simpa using hhx }, - obtain ⟨k, hk, hk'⟩ := IH _ _, - { refine ⟨k, _, _⟩, - { rw [←cycle_of_eq_one_iff] at hhx, - rwa [hd.cycle_of_mul_distrib, hhx, mul_one] }, - { simpa [←zpow_coe_nat, hpow] using hk' } }, - { use m, - simp [hpow] }, - { simpa [hhx] using hx } } } -end - -lemma same_cycle.nat [fintype α] (f : perm α) {x y : α} (h : same_cycle f x y) : - ∃ (i : ℕ) (hi : 0 < i) (hi' : i ≤ (f.cycle_of x).support.card + 1), (f ^ i) x = y := -begin - by_cases hx : x ∈ f.support, - { obtain ⟨k, hk, hk'⟩ := same_cycle.nat_of_mem_support f h hx, - cases k, - { refine ⟨(f.cycle_of x).support.card, _, self_le_add_right _ _, _⟩, - { refine zero_lt_one.trans (one_lt_card_support_of_ne_one _), - simpa using hx }, - { simp only [perm.coe_one, id.def, pow_zero] at hk', - subst hk', - rw [←(is_cycle_cycle_of _ $ mem_support.1 hx).order_of, - ←cycle_of_pow_apply_self, pow_order_of_eq_one, one_apply] } }, - { exact ⟨k + 1, by simp, nat.le_succ_of_le hk.le, hk'⟩ } }, - { refine ⟨1, zero_lt_one, by simp, _⟩, - obtain ⟨k, rfl⟩ := h, - rw [not_mem_support] at hx, - rw [pow_apply_eq_self_of_apply_eq_self hx, - zpow_apply_eq_self_of_apply_eq_self hx] } -end - section generation variables [finite β] diff --git a/src/group_theory/perm/cycle/concrete.lean b/src/group_theory/perm/cycle/concrete.lean index 1bc2a9449fa59..03a7eb22a2180 100644 --- a/src/group_theory/perm/cycle/concrete.lean +++ b/src/group_theory/perm/cycle/concrete.lean @@ -262,7 +262,7 @@ begin rw ←support_cycle_of_eq_nil_iff at hx, simp [hx] }, { rintro ⟨h, hx⟩, - simpa using same_cycle.nat_of_mem_support _ h hx } + simpa using h.exists_pow_eq_of_mem_support hx } end lemma nodup_to_list (p : perm α) (x : α) : @@ -305,7 +305,7 @@ lemma next_to_list_eq_apply (p : perm α) (x y : α) (hy : y ∈ to_list p x) : next (to_list p x) y hy = p y := begin rw mem_to_list_iff at hy, - obtain ⟨k, hk, hk'⟩ := hy.left.nat_of_mem_support _ hy.right, + obtain ⟨k, hk, hk'⟩ := hy.left.exists_pow_eq_of_mem_support hy.right, rw ←nth_le_to_list p x k (by simpa using hk) at hk', simp_rw ←hk', rw [next_nth_le _ (nodup_to_list _ _), nth_le_to_list, nth_le_to_list, ←mul_apply, ←pow_succ, @@ -317,7 +317,7 @@ lemma to_list_pow_apply_eq_rotate (p : perm α) (x : α) (k : ℕ) : p.to_list ((p ^ k) x) = (p.to_list x).rotate k := begin apply ext_le, - { simp }, + { simp only [length_to_list, cycle_of_self_apply_pow, length_rotate]}, { intros n hn hn', rw [nth_le_to_list, nth_le_rotate, nth_le_to_list, length_to_list, pow_mod_card_support_cycle_of_self_apply, pow_add, mul_apply] } @@ -327,7 +327,7 @@ lemma same_cycle.to_list_is_rotated {f : perm α} {x y : α} (h : same_cycle f x to_list f x ~r to_list f y := begin by_cases hx : x ∈ f.support, - { obtain ⟨_ | k, hk, hy⟩ := h.nat_of_mem_support _ hx, + { obtain ⟨_ | k, hk, hy⟩ := h.exists_pow_eq_of_mem_support hx, { simp only [coe_one, id.def, pow_zero] at hy, simp [hy] }, use k.succ, @@ -341,7 +341,7 @@ lemma pow_apply_mem_to_list_iff_mem_support {n : ℕ} : begin rw [mem_to_list_iff, and_iff_right_iff_imp], refine λ _, same_cycle.symm _, - rw same_cycle_pow_left_iff + rw same_cycle_pow_left end lemma to_list_form_perm_nil (x : α) : @@ -389,7 +389,7 @@ begin form_perm_nil] }, ext y, by_cases hy : same_cycle f x y, - { obtain ⟨k, hk, rfl⟩ := hy.nat_of_mem_support _ (mem_support.mpr hx), + { obtain ⟨k, hk, rfl⟩ := hy.exists_pow_eq_of_mem_support (mem_support.mpr hx), rw [cycle_of_apply_apply_pow_self, list.form_perm_apply_mem_eq_next (nodup_to_list f x), next_to_list_eq_apply, pow_succ, mul_apply], rw mem_to_list_iff, diff --git a/src/group_theory/perm/support.lean b/src/group_theory/perm/support.lean index b779f3197931f..f8aac0b2bf1e7 100644 --- a/src/group_theory/perm/support.lean +++ b/src/group_theory/perm/support.lean @@ -615,4 +615,8 @@ end card end support +@[simp] lemma support_subtype_perm [decidable_eq α] {s : finset α} (f : perm α) (h) : + (f.subtype_perm h : perm {x // x ∈ s}).support = s.attach.filter (λ x, f x ≠ x) := +by { ext, simp [subtype.ext_iff] } + end equiv.perm From 11c2b8c18d1a8e44fe9ba8ba6b931d51b4734150 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 10 Jan 2023 13:41:41 +0000 Subject: [PATCH 0198/1291] chore(data/real/ennreal): Protect ambiguous lemmas (#18076) Protect `ennreal.X` if declaration `X` already exists. --- src/analysis/special_functions/pow.lean | 7 +- src/data/real/ennreal.lean | 158 +++++++++--------- .../integral/mean_inequalities.lean | 4 +- src/order/filter/ennreal.lean | 4 +- src/topology/instances/ennreal.lean | 2 +- 5 files changed, 86 insertions(+), 89 deletions(-) diff --git a/src/analysis/special_functions/pow.lean b/src/analysis/special_functions/pow.lean index 448127aefa8fa..4f46b791aee68 100644 --- a/src/analysis/special_functions/pow.lean +++ b/src/analysis/special_functions/pow.lean @@ -1741,8 +1741,9 @@ begin replace hy := hy.lt_or_lt, rcases eq_or_ne x 0 with rfl|h0, { cases hy; simp * }, rcases eq_or_ne x ⊤ with rfl|h_top, { cases hy; simp * }, - apply eq_inv_of_mul_eq_one_left, - rw [← mul_rpow_of_ne_zero (inv_ne_zero.2 h_top) h0, inv_mul_cancel h0 h_top, one_rpow] + apply ennreal.eq_inv_of_mul_eq_one_left, + rw [← mul_rpow_of_ne_zero (ennreal.inv_ne_zero.2 h_top) h0, ennreal.inv_mul_cancel h0 h_top, + one_rpow] end lemma div_rpow_of_nonneg (x y : ℝ≥0∞) {z : ℝ} (hz : 0 ≤ z) : @@ -1874,7 +1875,7 @@ lemma rpow_pos {p : ℝ} {x : ℝ≥0∞} (hx_pos : 0 < x) (hx_ne_top : x ≠ begin cases lt_or_le 0 p with hp_pos hp_nonpos, { exact rpow_pos_of_nonneg hx_pos (le_of_lt hp_pos), }, - { rw [←neg_neg p, rpow_neg, inv_pos], + { rw [←neg_neg p, rpow_neg, ennreal.inv_pos], exact rpow_ne_top_of_nonneg (right.nonneg_neg_iff.mpr hp_nonpos) hx_ne_top, }, end diff --git a/src/data/real/ennreal.lean b/src/data/real/ennreal.lean index a46143ac853e7..d7b87d244de5d 100644 --- a/src/data/real/ennreal.lean +++ b/src/data/real/ennreal.lean @@ -1033,21 +1033,21 @@ begin rw [← coe_inv ha, ← coe_pow, ← coe_inv (pow_ne_zero _ ha), ← inv_pow, coe_pow] end -lemma mul_inv_cancel (h0 : a ≠ 0) (ht : a ≠ ∞) : a * a⁻¹ = 1 := +protected lemma mul_inv_cancel (h0 : a ≠ 0) (ht : a ≠ ∞) : a * a⁻¹ = 1 := begin lift a to ℝ≥0 using ht, norm_cast at *, exact mul_inv_cancel h0 end -lemma inv_mul_cancel (h0 : a ≠ 0) (ht : a ≠ ∞) : a⁻¹ * a = 1 := -mul_comm a a⁻¹ ▸ mul_inv_cancel h0 ht +protected lemma inv_mul_cancel (h0 : a ≠ 0) (ht : a ≠ ∞) : a⁻¹ * a = 1 := +mul_comm a a⁻¹ ▸ ennreal.mul_inv_cancel h0 ht -lemma div_mul_cancel (h0 : a ≠ 0) (hI : a ≠ ∞) : (b / a) * a = b := -by rw [div_eq_mul_inv, mul_assoc, inv_mul_cancel h0 hI, mul_one] +protected lemma div_mul_cancel (h0 : a ≠ 0) (hI : a ≠ ∞) : (b / a) * a = b := +by rw [div_eq_mul_inv, mul_assoc, ennreal.inv_mul_cancel h0 hI, mul_one] -lemma mul_div_cancel' (h0 : a ≠ 0) (hI : a ≠ ∞) : a * (b / a) = b := -by rw [mul_comm, div_mul_cancel h0 hI] +protected lemma mul_div_cancel' (h0 : a ≠ 0) (hI : a ≠ ∞) : a * (b / a) = b := +by rw [mul_comm, ennreal.div_mul_cancel h0 hI] instance : has_involutive_inv ℝ≥0∞ := { inv := has_inv.inv, @@ -1065,12 +1065,12 @@ by { simp only [lt_top_iff_ne_top, inv_ne_top, pos_iff_ne_zero] } lemma div_lt_top {x y : ℝ≥0∞} (h1 : x ≠ ∞) (h2 : y ≠ 0) : x / y < ∞ := mul_lt_top h1 (inv_ne_top.mpr h2) -@[simp] lemma inv_eq_zero : a⁻¹ = 0 ↔ a = ∞ := +@[simp] protected lemma inv_eq_zero : a⁻¹ = 0 ↔ a = ∞ := inv_top ▸ inv_inj -lemma inv_ne_zero : a⁻¹ ≠ 0 ↔ a ≠ ∞ := by simp +protected lemma inv_ne_zero : a⁻¹ ≠ 0 ↔ a ≠ ∞ := by simp -lemma mul_inv {a b : ℝ≥0∞} (ha : a ≠ 0 ∨ b ≠ ∞) (hb : a ≠ ∞ ∨ b ≠ 0) : +protected lemma mul_inv {a b : ℝ≥0∞} (ha : a ≠ 0 ∨ b ≠ ∞) (hb : a ≠ ∞ ∨ b ≠ 0) : (a * b)⁻¹ = a⁻¹ * b⁻¹ := begin induction b using with_top.rec_top_coe, @@ -1092,19 +1092,18 @@ end protected lemma mul_div_mul_left (a b : ℝ≥0∞) (hc : c ≠ 0) (hc' : c ≠ ⊤) : c * a / (c * b) = a / b := -by rw [div_eq_mul_inv, div_eq_mul_inv, mul_inv (or.inl hc) (or.inl hc'), mul_mul_mul_comm, +by rw [div_eq_mul_inv, div_eq_mul_inv, ennreal.mul_inv (or.inl hc) (or.inl hc'), mul_mul_mul_comm, ennreal.mul_inv_cancel hc hc', one_mul] protected lemma mul_div_mul_right (a b : ℝ≥0∞) (hc : c ≠ 0) (hc' : c ≠ ⊤) : a * c / (b * c) = a / b := -by rw [div_eq_mul_inv, div_eq_mul_inv, mul_inv (or.inr hc') (or.inr hc), mul_mul_mul_comm, +by rw [div_eq_mul_inv, div_eq_mul_inv, ennreal.mul_inv (or.inr hc') (or.inr hc), mul_mul_mul_comm, ennreal.mul_inv_cancel hc hc', mul_one] protected lemma sub_div (h : 0 < b → b < a → c ≠ 0) : (a - b) / c = a / c - b / c := by { simp_rw div_eq_mul_inv, exact ennreal.sub_mul (by simpa using h) } -@[simp] lemma inv_pos : 0 < a⁻¹ ↔ a ≠ ∞ := -pos_iff_ne_zero.trans inv_ne_zero +@[simp] protected lemma inv_pos : 0 < a⁻¹ ↔ a ≠ ∞ := pos_iff_ne_zero.trans ennreal.inv_ne_zero lemma inv_strict_anti : strict_anti (has_inv.inv : ℝ≥0∞ → ℝ≥0∞) := begin @@ -1117,31 +1116,27 @@ begin exact nnreal.inv_lt_inv ha h end -@[simp] lemma inv_lt_inv : a⁻¹ < b⁻¹ ↔ b < a := inv_strict_anti.lt_iff_lt +@[simp] protected lemma inv_lt_inv : a⁻¹ < b⁻¹ ↔ b < a := inv_strict_anti.lt_iff_lt lemma inv_lt_iff_inv_lt : a⁻¹ < b ↔ b⁻¹ < a := -by simpa only [inv_inv] using @inv_lt_inv a b⁻¹ +by simpa only [inv_inv] using @ennreal.inv_lt_inv a b⁻¹ lemma lt_inv_iff_lt_inv : a < b⁻¹ ↔ b < a⁻¹ := -by simpa only [inv_inv] using @inv_lt_inv a⁻¹ b +by simpa only [inv_inv] using @ennreal.inv_lt_inv a⁻¹ b @[simp, priority 1100] -- higher than le_inv_iff_mul_le -lemma inv_le_inv : a⁻¹ ≤ b⁻¹ ↔ b ≤ a := inv_strict_anti.le_iff_le +protected lemma inv_le_inv : a⁻¹ ≤ b⁻¹ ↔ b ≤ a := inv_strict_anti.le_iff_le lemma inv_le_iff_inv_le : a⁻¹ ≤ b ↔ b⁻¹ ≤ a := -by simpa only [inv_inv] using @inv_le_inv a b⁻¹ +by simpa only [inv_inv] using @ennreal.inv_le_inv a b⁻¹ lemma le_inv_iff_le_inv : a ≤ b⁻¹ ↔ b ≤ a⁻¹ := -by simpa only [inv_inv] using @inv_le_inv a⁻¹ b +by simpa only [inv_inv] using @ennreal.inv_le_inv a⁻¹ b -@[simp] lemma inv_le_one : a⁻¹ ≤ 1 ↔ 1 ≤ a := -inv_le_iff_inv_le.trans $ by rw inv_one - -lemma one_le_inv : 1 ≤ a⁻¹ ↔ a ≤ 1 := -le_inv_iff_le_inv.trans $ by rw inv_one - -@[simp] lemma inv_lt_one : a⁻¹ < 1 ↔ 1 < a := -inv_lt_iff_inv_lt.trans $ by rw [inv_one] +@[simp] protected lemma inv_le_one : a⁻¹ ≤ 1 ↔ 1 ≤ a := by rw [inv_le_iff_inv_le, inv_one] +protected lemma one_le_inv : 1 ≤ a⁻¹ ↔ a ≤ 1 := by rw [le_inv_iff_le_inv, inv_one] +@[simp] protected lemma inv_lt_one : a⁻¹ < 1 ↔ 1 < a := by rw [inv_lt_iff_inv_lt, inv_one] +@[simp] protected lemma one_lt_inv : 1 < a⁻¹ ↔ a < 1 := by rw [lt_inv_iff_lt_inv, inv_one] /-- The inverse map `λ x, x⁻¹` is an order isomorphism between `ℝ≥0∞` and its `order_dual` -/ @[simps apply] @@ -1153,10 +1148,10 @@ def _root_.order_iso.inv_ennreal : ℝ≥0∞ ≃o ℝ≥0∞ᵒᵈ := lemma _root_.order_iso.inv_ennreal_symm_apply : order_iso.inv_ennreal.symm a = (order_dual.of_dual a)⁻¹ := rfl -lemma pow_le_pow_of_le_one {n m : ℕ} (ha : a ≤ 1) (h : n ≤ m) : a ^ m ≤ a ^ n := +protected lemma pow_le_pow_of_le_one {n m : ℕ} (ha : a ≤ 1) (h : n ≤ m) : a ^ m ≤ a ^ n := begin - rw [←inv_inv a, ← ennreal.inv_pow, ← @ennreal.inv_pow a⁻¹, inv_le_inv], - exact pow_le_pow (one_le_inv.2 ha) h + rw [←inv_inv a, ← ennreal.inv_pow, ← @ennreal.inv_pow a⁻¹, ennreal.inv_le_inv], + exact pow_le_pow (ennreal.one_le_inv.2 ha) h end @[simp] lemma div_top : a / ∞ = 0 := by rw [div_eq_mul_inv, inv_top, mul_zero] @@ -1172,13 +1167,13 @@ top_div_of_ne_top h.ne lemma top_div : ∞ / a = if a = ∞ then 0 else ∞ := by by_cases a = ∞; simp [top_div_of_ne_top, *] -@[simp] lemma zero_div : 0 / a = 0 := zero_mul a⁻¹ +@[simp] protected lemma zero_div : 0 / a = 0 := zero_mul a⁻¹ lemma div_eq_top : a / b = ∞ ↔ (a ≠ 0 ∧ b = 0) ∨ (a = ∞ ∧ b ≠ ∞) := by simp [div_eq_mul_inv, ennreal.mul_eq_top] -lemma le_div_iff_mul_le (h0 : b ≠ 0 ∨ c ≠ 0) (ht : b ≠ ∞ ∨ c ≠ ∞) : - a ≤ c / b ↔ a * b ≤ c := +protected lemma le_div_iff_mul_le (h0 : b ≠ 0 ∨ c ≠ 0) (ht : b ≠ ∞ ∨ c ≠ ∞) : + a ≤ c / b ↔ a * b ≤ c := begin induction b using with_top.rec_top_coe, { lift c to ℝ≥0 using ht.neg_resolve_left rfl, @@ -1188,24 +1183,26 @@ begin { have hc : c ≠ 0, from h0.neg_resolve_left rfl, simp [div_zero hc] }, { rw [← coe_ne_zero] at hb, - rw [← ennreal.mul_le_mul_right hb coe_ne_top, div_mul_cancel hb coe_ne_top] }, + rw [← ennreal.mul_le_mul_right hb coe_ne_top, ennreal.div_mul_cancel hb coe_ne_top] }, end -lemma div_le_iff_le_mul (hb0 : b ≠ 0 ∨ c ≠ ∞) (hbt : b ≠ ∞ ∨ c ≠ 0) : a / b ≤ c ↔ a ≤ c * b := +protected lemma div_le_iff_le_mul (hb0 : b ≠ 0 ∨ c ≠ ∞) (hbt : b ≠ ∞ ∨ c ≠ 0) : + a / b ≤ c ↔ a ≤ c * b := begin suffices : a * b⁻¹ ≤ c ↔ a ≤ c / b⁻¹, by simpa [div_eq_mul_inv], - refine (le_div_iff_mul_le _ _).symm; simpa + refine (ennreal.le_div_iff_mul_le _ _).symm; simpa end -lemma lt_div_iff_mul_lt (hb0 : b ≠ 0 ∨ c ≠ ∞) (hbt : b ≠ ∞ ∨ c ≠ 0) : c < a / b ↔ c * b < a := -lt_iff_lt_of_le_iff_le (div_le_iff_le_mul hb0 hbt) +protected lemma lt_div_iff_mul_lt (hb0 : b ≠ 0 ∨ c ≠ ∞) (hbt : b ≠ ∞ ∨ c ≠ 0) : + c < a / b ↔ c * b < a := +lt_iff_lt_of_le_iff_le (ennreal.div_le_iff_le_mul hb0 hbt) lemma div_le_of_le_mul (h : a ≤ b * c) : a / c ≤ b := begin by_cases h0 : c = 0, { have : a = 0, by simpa [h0] using h, simp [*] }, by_cases hinf : c = ∞, by simp [hinf], - exact (div_le_iff_le_mul (or.inl h0) (or.inl hinf)).2 h + exact (ennreal.div_le_iff_le_mul (or.inl h0) (or.inl hinf)).2 h end lemma div_le_of_le_mul' (h : a ≤ b * c) : a / b ≤ c := @@ -1222,7 +1219,7 @@ mul_comm a c ▸ mul_le_of_le_div h protected lemma div_lt_iff (h0 : b ≠ 0 ∨ c ≠ 0) (ht : b ≠ ∞ ∨ c ≠ ∞) : c / b < a ↔ c < a * b := -lt_iff_lt_of_le_iff_le $ le_div_iff_mul_le h0 ht +lt_iff_lt_of_le_iff_le $ ennreal.le_div_iff_mul_le h0 ht lemma mul_lt_of_lt_div (h : a < b / c) : a * c < b := by { contrapose! h, exact ennreal.div_le_of_le_mul h } @@ -1231,14 +1228,14 @@ lemma mul_lt_of_lt_div' (h : a < b / c) : c * a < b := mul_comm a c ▸ mul_lt_o lemma inv_le_iff_le_mul (h₁ : b = ∞ → a ≠ 0) (h₂ : a = ∞ → b ≠ 0) : a⁻¹ ≤ b ↔ 1 ≤ a * b := begin - rw [← one_div, div_le_iff_le_mul, mul_comm], + rw [← one_div, ennreal.div_le_iff_le_mul, mul_comm], exacts [or_not_of_imp h₁, not_or_of_imp h₂] end @[simp] lemma le_inv_iff_mul_le : a ≤ b⁻¹ ↔ a * b ≤ 1 := -by rw [← one_div, le_div_iff_mul_le]; { right, simp } +by rw [← one_div, ennreal.le_div_iff_mul_le]; { right, simp } -lemma div_le_div {a b c d : ℝ≥0∞} (hab : a ≤ b) (hdc : d ≤ c) : a / c ≤ b / d := +protected lemma div_le_div (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) protected lemma div_le_div_left (h : a ≤ b) (c : ℝ≥0∞) : c / b ≤ c / a := @@ -1247,16 +1244,15 @@ ennreal.div_le_div le_rfl h protected lemma div_le_div_right (h : a ≤ b) (c : ℝ≥0∞) : a / c ≤ b / c := ennreal.div_le_div h le_rfl -lemma eq_inv_of_mul_eq_one_left (h : a * b = 1) : a = b⁻¹ := +protected lemma eq_inv_of_mul_eq_one_left (h : a * b = 1) : a = b⁻¹ := begin - have hb : b ≠ ∞, - { rintro rfl, - simpa [left_ne_zero_of_mul_eq_one h] using h }, - rw [← mul_one a, ← mul_inv_cancel (right_ne_zero_of_mul_eq_one h) hb, ← mul_assoc, h, one_mul] + rw [←mul_one a, ←ennreal.mul_inv_cancel (right_ne_zero_of_mul_eq_one h), ←mul_assoc, h, one_mul], + rintro rfl, + simpa [left_ne_zero_of_mul_eq_one h] using h, end lemma mul_le_iff_le_inv {a b r : ℝ≥0∞} (hr₀ : r ≠ 0) (hr₁ : r ≠ ∞) : (r * a ≤ b ↔ a ≤ r⁻¹ * b) := -by rw [← @ennreal.mul_le_mul_left _ a _ hr₀ hr₁, ← mul_assoc, mul_inv_cancel hr₀ hr₁, one_mul] +by rw [←@ennreal.mul_le_mul_left _ a _ hr₀ hr₁, ←mul_assoc, ennreal.mul_inv_cancel hr₀ hr₁, one_mul] lemma le_of_forall_nnreal_lt {x y : ℝ≥0∞} (h : ∀ r : ℝ≥0, ↑r < x → ↑r ≤ y) : x ≤ y := begin @@ -1271,23 +1267,23 @@ le_of_forall_nnreal_lt $ λ r hr, (zero_le r).eq_or_lt.elim (λ h, h ▸ zero_le lemma eq_top_of_forall_nnreal_le {x : ℝ≥0∞} (h : ∀ r : ℝ≥0, ↑r ≤ x) : x = ∞ := top_unique $ le_of_forall_nnreal_lt $ λ r hr, h r -lemma add_div : (a + b) / c = a / c + b / c := right_distrib a b (c⁻¹) +protected lemma add_div : (a + b) / c = a / c + b / c := right_distrib a b (c⁻¹) -lemma div_add_div_same {a b c : ℝ≥0∞} : a / c + b / c = (a + b) / c := -add_div.symm +protected lemma div_add_div_same {a b c : ℝ≥0∞} : a / c + b / c = (a + b) / c := +ennreal.add_div.symm -lemma div_self (h0 : a ≠ 0) (hI : a ≠ ∞) : a / a = 1 := -mul_inv_cancel h0 hI +protected lemma div_self (h0 : a ≠ 0) (hI : a ≠ ∞) : a / a = 1 := +ennreal.mul_inv_cancel h0 hI lemma mul_div_le : a * (b / a) ≤ b := mul_le_of_le_div' le_rfl -- TODO: add this lemma for an `is_unit` in any `division_monoid` lemma eq_div_iff (ha : a ≠ 0) (ha' : a ≠ ∞) : b = c / a ↔ a * b = c := -⟨λ h, by rw [h, mul_div_cancel' ha ha'], - λ h, by rw [← h, mul_div_assoc, mul_div_cancel' ha ha']⟩ +⟨λ h, by rw [h, ennreal.mul_div_cancel' ha ha'], + λ h, by rw [← h, mul_div_assoc, ennreal.mul_div_cancel' ha ha']⟩ -lemma div_eq_div_iff (ha : a ≠ 0) (ha' : a ≠ ∞) (hb : b ≠ 0) (hb' : b ≠ ∞) : +protected lemma div_eq_div_iff (ha : a ≠ 0) (ha' : a ≠ ∞) (hb : b ≠ 0) (hb' : b ≠ ∞) : c / b = d / a ↔ a * c = b * d := begin rw eq_div_iff ha ha', @@ -1296,10 +1292,10 @@ begin end lemma div_eq_one_iff {a b : ℝ≥0∞} (hb₀ : b ≠ 0) (hb₁ : b ≠ ∞) : a / b = 1 ↔ a = b := -⟨λ h, by rw [← (eq_div_iff hb₀ hb₁).mp h.symm, mul_one], λ h, h.symm ▸ div_self hb₀ hb₁⟩ +⟨λ h, by rw [← (eq_div_iff hb₀ hb₁).mp h.symm, mul_one], λ h, h.symm ▸ ennreal.div_self hb₀ hb₁⟩ lemma inv_two_add_inv_two : (2:ℝ≥0∞)⁻¹ + 2⁻¹ = 1 := -by rw [← two_mul, ← div_eq_mul_inv, div_self two_ne_zero two_ne_top] +by rw [← two_mul, ← div_eq_mul_inv, ennreal.div_self two_ne_zero two_ne_top] lemma inv_three_add_inv_three : (3 : ℝ≥0∞)⁻¹ + 3⁻¹ + 3⁻¹ = 1 := begin @@ -1308,7 +1304,7 @@ begin end @[simp] -lemma add_halves (a : ℝ≥0∞) : a / 2 + a / 2 = a := +protected lemma add_halves (a : ℝ≥0∞) : a / 2 + a / 2 = a := by rw [div_eq_mul_inv, ← mul_add, inv_two_add_inv_two, mul_one] @[simp] @@ -1321,12 +1317,11 @@ by simp [div_eq_mul_inv] @[simp] lemma div_pos_iff : 0 < a / b ↔ a ≠ 0 ∧ b ≠ ∞ := by simp [pos_iff_ne_zero, not_or_distrib] -lemma half_pos {a : ℝ≥0∞} (h : a ≠ 0) : 0 < a / 2 := -by simp [h] +protected lemma half_pos (h : a ≠ 0) : 0 < a / 2 := by simp [h] -lemma one_half_lt_one : (2⁻¹ : ℝ≥0∞) < 1 := inv_lt_one.2 $ one_lt_two +protected lemma one_half_lt_one : (2⁻¹ : ℝ≥0∞) < 1 := ennreal.inv_lt_one.2 $ one_lt_two -lemma half_lt_self {a : ℝ≥0∞} (hz : a ≠ 0) (ht : a ≠ ∞) : a / 2 < a := +protected lemma half_lt_self (hz : a ≠ 0) (ht : a ≠ ∞) : a / 2 < a := begin lift a to ℝ≥0 using ht, rw coe_ne_zero at hz, @@ -1334,12 +1329,12 @@ begin exacts [nnreal.half_lt_self hz, two_ne_zero' _] end -lemma half_le_self : a / 2 ≤ a := le_add_self.trans_eq (add_halves _) +protected lemma half_le_self : a / 2 ≤ a := le_add_self.trans_eq $ ennreal.add_halves _ lemma sub_half (h : a ≠ ∞) : a - a / 2 = a / 2 := begin lift a to ℝ≥0 using h, - exact sub_eq_of_add_eq (mul_ne_top coe_ne_top $ by simp) (add_halves a) + exact sub_eq_of_add_eq (mul_ne_top coe_ne_top $ by simp) (ennreal.add_halves a) end @[simp] lemma one_sub_inv_two : (1:ℝ≥0∞) - 2⁻¹ = 2⁻¹ := @@ -1348,10 +1343,11 @@ by simpa only [div_eq_mul_inv, one_mul] using sub_half one_ne_top /-- The birational order isomorphism between `ℝ≥0∞` and the unit interval `set.Iic (1 : ℝ≥0∞)`. -/ @[simps apply_coe] def order_iso_Iic_one_birational : ℝ≥0∞ ≃o Iic (1 : ℝ≥0∞) := begin - refine strict_mono.order_iso_of_right_inverse (λ x, ⟨(x⁻¹ + 1)⁻¹, inv_le_one.2 $ le_add_self⟩) - (λ x y hxy, _) (λ x, (x⁻¹ - 1)⁻¹) (λ x, subtype.ext _), - { simpa only [subtype.mk_lt_mk, inv_lt_inv, ennreal.add_lt_add_iff_right one_ne_top] }, - { have : (1 : ℝ≥0∞) ≤ x⁻¹, from one_le_inv.2 x.2, + refine strict_mono.order_iso_of_right_inverse + (λ x, ⟨(x⁻¹ + 1)⁻¹, ennreal.inv_le_one.2 $ le_add_self⟩) (λ x y hxy, _) (λ x, (x⁻¹ - 1)⁻¹) + (λ x, subtype.ext _), + { simpa only [subtype.mk_lt_mk, ennreal.inv_lt_inv, ennreal.add_lt_add_iff_right one_ne_top] }, + { have : (1 : ℝ≥0∞) ≤ x⁻¹, from ennreal.one_le_inv.2 x.2, simp only [inv_inv, subtype.coe_mk, tsub_add_cancel_of_le this] } end @@ -1383,7 +1379,7 @@ rfl lemma exists_inv_nat_lt {a : ℝ≥0∞} (h : a ≠ 0) : ∃n:ℕ, (n:ℝ≥0∞)⁻¹ < a := -inv_inv a ▸ by simp only [inv_lt_inv, ennreal.exists_nat_gt (inv_ne_top.2 h)] +inv_inv a ▸ by simp only [ennreal.inv_lt_inv, ennreal.exists_nat_gt (inv_ne_top.2 h)] lemma exists_nat_pos_mul_gt (ha : a ≠ 0) (hb : b ≠ ∞) : ∃ n > 0, b < (n : ℕ) * a := @@ -1405,8 +1401,8 @@ begin rcases exists_nat_pos_mul_gt hb ha with ⟨n, npos, hn⟩, have : (n : ℝ≥0∞) ≠ 0 := nat.cast_ne_zero.2 npos.lt.ne', use [n, npos], - rwa [← one_mul b, ← inv_mul_cancel this (nat_ne_top n), - mul_assoc, mul_lt_mul_left (inv_ne_zero.2 $ nat_ne_top _) (inv_ne_top.2 this)] + rwa [← one_mul b, ← ennreal.inv_mul_cancel this (nat_ne_top n), mul_assoc, + mul_lt_mul_left (ennreal.inv_ne_zero.2 $ nat_ne_top _) (inv_ne_top.2 this)] end lemma exists_nnreal_pos_mul_lt (ha : a ≠ ∞) (hb : b ≠ 0) : @@ -1422,7 +1418,7 @@ lemma exists_inv_two_pow_lt (ha : a ≠ 0) : begin rcases exists_inv_nat_lt ha with ⟨n, hn⟩, refine ⟨n, lt_trans _ hn⟩, - rw [← ennreal.inv_pow, inv_lt_inv], + rw [← ennreal.inv_pow, ennreal.inv_lt_inv], norm_cast, exact n.lt_two_pow end @@ -1440,7 +1436,7 @@ begin cases n, { exact ennreal.pow_pos ha.bot_lt n }, { simp only [h'a, pow_eq_top_iff, zpow_neg_succ_of_nat, ne.def, not_false_iff, - inv_pos, false_and] } + ennreal.inv_pos, false_and] } end lemma zpow_lt_top (ha : a ≠ 0) (h'a : a ≠ ∞) (n : ℤ) : a ^ n < ∞ := @@ -1504,9 +1500,9 @@ begin { apply absurd h (not_le_of_gt _), exact lt_of_lt_of_le (int.neg_succ_lt_zero _) (int.of_nat_nonneg _) }, { simp only [zpow_neg_succ_of_nat, int.of_nat_eq_coe, zpow_coe_nat], - refine le_trans (inv_le_one.2 _) _; + refine (ennreal.inv_le_one.2 _).trans _; exact ennreal.one_le_pow_of_one_le hx _, }, - { simp only [zpow_neg_succ_of_nat, inv_le_inv], + { simp only [zpow_neg_succ_of_nat, ennreal.inv_le_inv], apply pow_le_pow hx, simpa only [←int.coe_nat_le_coe_nat_iff, neg_le_neg_iff, int.coe_nat_add, int.coe_nat_one, int.neg_succ_of_nat_eq] using h } @@ -1515,7 +1511,7 @@ end lemma monotone_zpow {x : ℝ≥0∞} (hx : 1 ≤ x) : monotone ((^) x : ℤ → ℝ≥0∞) := λ a b h, zpow_le_of_le hx h -lemma zpow_add {x : ℝ≥0∞} (hx : x ≠ 0) (h'x : x ≠ ∞) (m n : ℤ) : +protected lemma zpow_add {x : ℝ≥0∞} (hx : x ≠ 0) (h'x : x ≠ ∞) (m n : ℤ) : x ^ (m + n) = x ^ m * x ^ n := begin lift x to ℝ≥0 using h'x, @@ -1925,8 +1921,8 @@ lemma infi_mul_of_ne {ι} {f : ι → ℝ≥0∞} {x : ℝ≥0∞} (h0 : x ≠ 0 infi f * x = ⨅ i, f i * x := le_antisymm mul_right_mono.map_infi_le - ((div_le_iff_le_mul (or.inl h0) $ or.inl h).mp $ le_infi $ - λ i, (div_le_iff_le_mul (or.inl h0) $ or.inl h).mpr $ infi_le _ _) + ((ennreal.div_le_iff_le_mul (or.inl h0) $ or.inl h).mp $ le_infi $ + λ i, (ennreal.div_le_iff_le_mul (or.inl h0) $ or.inl h).mpr $ infi_le _ _) /-- If `x ≠ ∞`, then right multiplication by `x` maps infimum over a nonempty type to infimum. See also `ennreal.infi_mul_of_ne` that assumes `x ≠ 0` but does not require `[nonempty ι]`. -/ diff --git a/src/measure_theory/integral/mean_inequalities.lean b/src/measure_theory/integral/mean_inequalities.lean index 614e0060f3520..e323bac178bdf 100644 --- a/src/measure_theory/integral/mean_inequalities.lean +++ b/src/measure_theory/integral/mean_inequalities.lean @@ -74,7 +74,7 @@ def fun_mul_inv_snorm (f : α → ℝ≥0∞) (p : ℝ) (μ : measure α) : α lemma fun_eq_fun_mul_inv_snorm_mul_snorm {p : ℝ} (f : α → ℝ≥0∞) (hf_nonzero : ∫⁻ a, (f a) ^ p ∂μ ≠ 0) (hf_top : ∫⁻ a, (f a) ^ p ∂μ ≠ ⊤) {a : α} : f a = (fun_mul_inv_snorm f p μ a) * (∫⁻ c, (f c)^p ∂μ)^(1/p) := -by simp [fun_mul_inv_snorm, mul_assoc, inv_mul_cancel, hf_nonzero, hf_top] +by simp [fun_mul_inv_snorm, mul_assoc, ennreal.inv_mul_cancel, hf_nonzero, hf_top] lemma fun_mul_inv_snorm_rpow {p : ℝ} (hp0 : 0 < p) {f : α → ℝ≥0∞} {a : α} : (fun_mul_inv_snorm f p μ a) ^ p = (f a)^p * (∫⁻ c, (f c) ^ p ∂μ)⁻¹ := @@ -90,7 +90,7 @@ lemma lintegral_rpow_fun_mul_inv_snorm_eq_one {p : ℝ} (hp0_lt : 0 < p) {f : α ∫⁻ c, (fun_mul_inv_snorm f p μ c)^p ∂μ = 1 := begin simp_rw fun_mul_inv_snorm_rpow hp0_lt, - rw [lintegral_mul_const', mul_inv_cancel hf_nonzero hf_top], + rw [lintegral_mul_const', ennreal.mul_inv_cancel hf_nonzero hf_top], rwa inv_ne_top end diff --git a/src/order/filter/ennreal.lean b/src/order/filter/ennreal.lean index a061290a9960b..580c082a1d442 100644 --- a/src/order/filter/ennreal.lean +++ b/src/order/filter/ennreal.lean @@ -60,8 +60,8 @@ begin let g := λ x : ℝ≥0∞, a * x, have hg_bij : function.bijective g, from function.bijective_iff_has_inverse.mpr ⟨(λ x, a⁻¹ * x), - ⟨λ x, by simp [←mul_assoc, inv_mul_cancel ha_zero ha_top], - λ x, by simp [g, ←mul_assoc, mul_inv_cancel ha_zero ha_top]⟩⟩, + ⟨λ x, by simp [←mul_assoc, ennreal.inv_mul_cancel ha_zero ha_top], + λ x, by simp [g, ←mul_assoc, ennreal.mul_inv_cancel ha_zero ha_top]⟩⟩, have hg_mono : strict_mono g, from monotone.strict_mono_of_injective (λ _ _ _, by rwa mul_le_mul_left ha_zero ha_top) hg_bij.1, diff --git a/src/topology/instances/ennreal.lean b/src/topology/instances/ennreal.lean index dfe2b9fc47fd8..9d679d95c2d7c 100644 --- a/src/topology/instances/ennreal.lean +++ b/src/topology/instances/ennreal.lean @@ -302,7 +302,7 @@ begin have : ∀ᶠ c : ℝ≥0∞ × ℝ≥0∞ in 𝓝 (∞, b), ↑n / ↑ε < c.1 ∧ ↑ε < c.2, from (lt_mem_nhds $ div_lt_top coe_ne_top hε.ne').prod_nhds (lt_mem_nhds hεb), refine this.mono (λ c hc, _), - exact (div_mul_cancel hε.ne' coe_ne_top).symm.trans_lt (mul_lt_mul hc.1 hc.2) + exact (ennreal.div_mul_cancel hε.ne' coe_ne_top).symm.trans_lt (mul_lt_mul hc.1 hc.2) end, begin cases a, {simp [none_eq_top] at hb, simp [none_eq_top, ht b hb, top_mul, hb] }, From 65902a4a1a39ff6fdfb657a35dde2579dbb4a155 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 10 Jan 2023 13:41:42 +0000 Subject: [PATCH 0199/1291] =?UTF-8?q?feat(data/mv=5Fpolynomial/*):=20Simpl?= =?UTF-8?q?e=20`=E2=80=A2`=20lemmas=20(#18084)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit A handful of missing lemmas about `a • f` where `f` is a multivariate polynomial. --- src/data/mv_polynomial/basic.lean | 7 +++++++ src/data/mv_polynomial/variables.lean | 5 +++++ 2 files changed, 12 insertions(+) diff --git a/src/data/mv_polynomial/basic.lean b/src/data/mv_polynomial/basic.lean index 691eeed210ef4..a361af8d54c34 100644 --- a/src/data/mv_polynomial/basic.lean +++ b/src/data/mv_polynomial/basic.lean @@ -419,6 +419,10 @@ by rw [X_pow_eq_monomial, support_monomial, if_neg (one_ne_zero' R)] @[simp] lemma support_zero : (0 : mv_polynomial σ R).support = ∅ := rfl +lemma support_smul [distrib_mul_action R S₁] {a : R} {f : mv_polynomial σ S₁} : + (a • f).support ⊆ f.support := +finsupp.support_smul + lemma support_sum {α : Type*} {s : finset α} {f : α → mv_polynomial σ R} : (∑ x in s, f x).support ⊆ s.bUnion (λ x, (f x).support) := finsupp.support_finset_sum @@ -668,6 +672,9 @@ variables (R) by simp [constant_coeff_eq] variables {R} +@[simp] lemma constant_coeff_smul [distrib_mul_action R S₁] (a : R) (f : mv_polynomial σ S₁) : + constant_coeff (a • f) = a • constant_coeff f := rfl + lemma constant_coeff_monomial [decidable_eq σ] (d : σ →₀ ℕ) (r : R) : constant_coeff (monomial d r) = if d = 0 then r else 0 := by rw [constant_coeff_eq, coeff_monomial] diff --git a/src/data/mv_polynomial/variables.lean b/src/data/mv_polynomial/variables.lean index 5a32886dc6843..b7ce2c7c96b80 100644 --- a/src/data/mv_polynomial/variables.lean +++ b/src/data/mv_polynomial/variables.lean @@ -602,6 +602,11 @@ finset.sup_le $ assume n hn, { assume a b₁ b₂, refl } end +lemma total_degree_smul_le [comm_semiring S] [distrib_mul_action R S] (a : R) + (f : mv_polynomial σ S) : + (a • f).total_degree ≤ f.total_degree := +finset.sup_mono support_smul + lemma total_degree_pow (a : mv_polynomial σ R) (n : ℕ) : (a ^ n).total_degree ≤ n * a.total_degree := begin From a2d2e18906e2b62627646b5d5be856e6a642062f Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9mi=20Bottinelli?= Date: Tue, 10 Jan 2023 18:13:13 +0000 Subject: [PATCH 0200/1291] feat(analysis/bounded_variation): some lemmas (#18108) * The variation of a function on a set is zero iff the image of the set has zero diameter. * Two functions that are pointwise at distance zero on a set have equal variation on that set. --- src/analysis/bounded_variation.lean | 63 ++++++++++++-------- src/topology/metric_space/emetric_space.lean | 13 ++++ 2 files changed, 52 insertions(+), 24 deletions(-) diff --git a/src/analysis/bounded_variation.lean b/src/analysis/bounded_variation.lean index 9adb29ee3322c..bb435c8f434af 100644 --- a/src/analysis/bounded_variation.lean +++ b/src/analysis/bounded_variation.lean @@ -75,16 +75,19 @@ begin exact ⟨⟨λ i, x, λ i j hij, le_rfl, λ i, hx⟩⟩, end -lemma eq_of_eq_on {f f' : α → E} {s : set α} (h : set.eq_on f f' s) : +lemma eq_of_edist_zero_on {f f' : α → E} {s : set α} (h : ∀ ⦃x⦄, x ∈ s → edist (f x) (f' x) = 0) : evariation_on f s = evariation_on f' s := begin dsimp only [evariation_on], congr' 1 with p : 1, congr' 1 with i : 1, - congr' 1; - exact h (p.2.2.2 _), + rw [edist_congr_right (h $ p.snd.prop.2 (i+1)), edist_congr_left (h $ p.snd.prop.2 i)], end +lemma eq_of_eq_on {f f' : α → E} {s : set α} (h : set.eq_on f f' s) : + evariation_on f s = evariation_on f' s := +eq_of_edist_zero_on (λ x xs, by rw [h xs, edist_self]) + lemma sum_le (f : α → E) {s : set α} (n : ℕ) {u : ℕ → α} (hu : monotone u) (us : ∀ i, u i ∈ s) : ∑ i in finset.range n, edist (f (u (i+1))) (f (u i)) ≤ evariation_on f s := @@ -171,19 +174,6 @@ lemma _root_.has_bounded_variation_on.has_locally_bounded_variation_on {f : α (h : has_bounded_variation_on f s) : has_locally_bounded_variation_on f s := λ x y hx hy, h.mono (inter_subset_left _ _) -lemma constant_on {f : α → E} {s : set α} - (hf : (f '' s).subsingleton) : evariation_on f s = 0 := -begin - apply le_antisymm _ (zero_le _), - apply supr_le _, - rintros ⟨n, ⟨u, hu, ut⟩⟩, - have : ∀ i, f (u i) = f (u 0) := λ i, hf ⟨u i, ut i, rfl⟩ ⟨u 0, ut 0, rfl⟩, - simp [subtype.coe_mk, le_zero_iff, finset.sum_eq_zero_iff, finset.mem_range, this], -end - -@[simp] protected lemma subsingleton (f : α → E) {s : set α} (hs : s.subsingleton) : - evariation_on f s = 0 := constant_on (hs.image f) - lemma edist_le (f : α → E) {s : set α} {x y : α} (hx : x ∈ s) (hy : y ∈ s) : edist (f x) (f y) ≤ evariation_on f s := begin @@ -206,6 +196,30 @@ begin simp [u, edist_comm], end +lemma eq_zero_iff (f : α → E) {s : set α} : + evariation_on f s = 0 ↔ ∀ (x y ∈ s), edist (f x) (f y) = 0 := +begin + split, + { rintro h x xs y ys, + rw [←le_zero_iff, ←h], + exact edist_le f xs ys, }, + { rintro h, + dsimp only [evariation_on], + rw ennreal.supr_eq_zero, + rintro ⟨n, u, um, us⟩, + exact finset.sum_eq_zero (λ i hi, h _ (us i.succ) _ (us i)), }, +end + +lemma constant_on {f : α → E} {s : set α} (hf : (f '' s).subsingleton) : evariation_on f s = 0 := +begin + rw eq_zero_iff, + rintro x xs y ys, + rw [hf ⟨x, xs, rfl⟩ ⟨y, ys, rfl⟩, edist_self], +end + +@[simp] protected lemma subsingleton (f : α → E) {s : set α} (hs : s.subsingleton) : + evariation_on f s = 0 := constant_on (hs.image f) + lemma lower_continuous_aux {ι : Type*} {F : ι → α → E} {p : filter ι} {f : α → E} {s : set α} (Ffs : ∀ x ∈ s, tendsto (λ i, F i x) p (𝓝 (f x))) {v : ℝ≥0∞} (hv : v < evariation_on f s) : ∀ᶠ (n : ι) in p, v < evariation_on (F n) s := @@ -438,9 +452,9 @@ begin begin rw [finset.sum_Ico_consecutive, finset.sum_Ico_consecutive, finset.range_eq_Ico], { exact zero_le _ }, - { linarith }, + { exact nat.succ_le_succ hN.left }, { exact zero_le _ }, - { linarith } + { exact N.pred_le.trans (N.le_succ) } end } end @@ -474,7 +488,7 @@ begin split_ifs, { exact hu hij }, { apply h _ (us _) _ (vt _) }, - { linarith }, + { exfalso, exact h_1 (hij.trans h_2), }, { apply hv (tsub_le_tsub hij le_rfl) } }, calc ∑ i in finset.range n, edist (f (u (i + 1))) (f (u i)) + ∑ (i : ℕ) in finset.range m, edist (f (v (i + 1))) (f (v i)) @@ -485,16 +499,16 @@ begin congr' 1, { apply finset.sum_congr rfl (λ i hi, _), simp only [finset.mem_range] at hi, - have : i + 1 ≤ n, by linarith, + have : i + 1 ≤ n := nat.succ_le_of_lt hi, simp [hi.le, this] }, { apply finset.sum_congr rfl (λ i hi, _), simp only [finset.mem_range] at hi, - have A : ¬(n + 1 + i + 1 ≤ n), by linarith, have B : ¬(n + 1 + i ≤ n), by linarith, + have A : ¬(n + 1 + i + 1 ≤ n) := λ h, B ((n+1+i).le_succ.trans h), have C : n + 1 + i - n = i + 1, { rw tsub_eq_iff_eq_add_of_le, { abel }, - { linarith } }, + { exact n.le_succ.trans (n.succ.le_add_right i), } }, simp only [A, B, C, nat.succ_sub_succ_eq_sub, if_false, add_tsub_cancel_left] } end ... = ∑ i in finset.range n, edist (f (w (i + 1))) (f (w i)) @@ -512,11 +526,11 @@ begin rintros i hi, simp only [finset.mem_union, finset.mem_range, finset.mem_Ico] at hi ⊢, cases hi, - { linarith }, + { exact lt_of_lt_of_le hi (n.le_succ.trans (n.succ.le_add_right m)) }, { exact hi.2 } }, { apply finset.disjoint_left.2 (λ i hi h'i, _), simp only [finset.mem_Ico, finset.mem_range] at hi h'i, - linarith [h'i.1] } + exact hi.not_lt (nat.lt_of_succ_le h'i.left) } end ... ≤ evariation_on f (s ∪ t) : sum_le f _ hw wst end @@ -871,3 +885,4 @@ lemma lipschitz_with.ae_differentiable_at {C : ℝ≥0} {f : ℝ → V} (h : lipschitz_with C f) : ∀ᵐ x, differentiable_at ℝ f x := (h.has_locally_bounded_variation_on univ).ae_differentiable_at + diff --git a/src/topology/metric_space/emetric_space.lean b/src/topology/metric_space/emetric_space.lean index a1e8643c8967b..1146f5df46403 100644 --- a/src/topology/metric_space/emetric_space.lean +++ b/src/topology/metric_space/emetric_space.lean @@ -115,6 +115,19 @@ by rw edist_comm z; apply edist_triangle theorem edist_triangle_right (x y z : α) : edist x y ≤ edist x z + edist y z := by rw edist_comm y; apply edist_triangle +lemma edist_congr_right {x y z : α} (h : edist x y = 0) : edist x z = edist y z := +begin + apply le_antisymm, + { rw [←zero_add (edist y z), ←h], + apply edist_triangle, }, + { rw edist_comm at h, + rw [←zero_add (edist x z), ←h], + apply edist_triangle, }, +end + +lemma edist_congr_left {x y z : α} (h : edist x y = 0) : edist z x = edist z y := +by { rw [edist_comm z x, edist_comm z y], apply edist_congr_right h, } + lemma edist_triangle4 (x y z t : α) : edist x t ≤ edist x y + edist y z + edist z t := calc From 6133ae2da6ae6693248bb5451de703f1ef154cc8 Mon Sep 17 00:00:00 2001 From: Chris Hughes Date: Tue, 10 Jan 2023 20:13:33 +0000 Subject: [PATCH 0201/1291] chore: remove uses of and.rec (#18123) --- src/data/complex/basic.lean | 2 +- src/data/fin/tuple/basic.lean | 2 +- src/data/list/dedup.lean | 2 +- 3 files changed, 3 insertions(+), 3 deletions(-) diff --git a/src/data/complex/basic.lean b/src/data/complex/basic.lean index b096ec98f6e83..91a575a64b484 100644 --- a/src/data/complex/basic.lean +++ b/src/data/complex/basic.lean @@ -46,7 +46,7 @@ theorem ext : ∀ {z w : ℂ}, z.re = w.re → z.im = w.im → z = w | ⟨zr, zi⟩ ⟨_, _⟩ rfl rfl := rfl theorem ext_iff {z w : ℂ} : z = w ↔ z.re = w.re ∧ z.im = w.im := -⟨λ H, by simp [H], and.rec ext⟩ +⟨λ H, by simp [H], λ h, ext h.1 h.2⟩ theorem re_surjective : surjective re := λ x, ⟨⟨x, 0⟩, rfl⟩ theorem im_surjective : surjective im := λ y, ⟨⟨0, y⟩, rfl⟩ diff --git a/src/data/fin/tuple/basic.lean b/src/data/fin/tuple/basic.lean index 63dbee10e3c9d..290eab4b23d51 100644 --- a/src/data/fin/tuple/basic.lean +++ b/src/data/fin/tuple/basic.lean @@ -165,7 +165,7 @@ end lemma cons_injective_iff {α} {x₀ : α} {x : fin n → α} : function.injective (cons x₀ x : fin n.succ → α) ↔ x₀ ∉ set.range x ∧ function.injective x := begin - refine ⟨λ h, ⟨_, _⟩, and.rec cons_injective_of_injective⟩, + refine ⟨λ h, ⟨_, _⟩, λ h, cons_injective_of_injective h.1 h.2⟩, { rintros ⟨i, hi⟩, replace h := @h i.succ 0, simpa [hi, succ_ne_zero] using h, }, diff --git a/src/data/list/dedup.lean b/src/data/list/dedup.lean index 20dfcd565439e..a0a7ee3667055 100644 --- a/src/data/list/dedup.lean +++ b/src/data/list/dedup.lean @@ -35,7 +35,7 @@ pw_filter_cons_of_pos $ by simpa only [forall_mem_ne] using h @[simp] theorem mem_dedup {a : α} {l : list α} : a ∈ dedup l ↔ a ∈ l := by simpa only [dedup, forall_mem_ne, not_not] using not_congr (@forall_mem_pw_filter α (≠) _ - (λ x y z xz, not_and_distrib.1 $ mt (and.rec eq.trans) xz) a l) + (λ x y z xz, not_and_distrib.1 $ mt (λ h, eq.trans h.1 h.2) xz) a l) @[simp] theorem dedup_cons_of_mem {a : α} {l : list α} (h : a ∈ l) : dedup (a :: l) = dedup l := From e54dc27b9a864c9c59de93953ba6b47bfc909ee3 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Tue, 10 Jan 2023 23:11:36 +0000 Subject: [PATCH 0202/1291] feat(order/filter/pointwise): add lemmas (#18098) --- src/order/filter/pointwise.lean | 16 +++++++++++++++- 1 file changed, 15 insertions(+), 1 deletion(-) diff --git a/src/order/filter/pointwise.lean b/src/order/filter/pointwise.lean index 0759bcb0ef67f..a1b5b63a54a6e 100644 --- a/src/order/filter/pointwise.lean +++ b/src/order/filter/pointwise.lean @@ -84,6 +84,9 @@ localized "attribute [instance] filter.has_one filter.has_zero" in pointwise tendsto f a 1 ↔ ∀ᶠ x in a, f x = 1 := tendsto_pure +@[simp, to_additive] lemma one_prod_one [has_one β] : (1 : filter α) ×ᶠ (1 : filter β) = 1 := +prod_pure_pure + /-- `pure` as a `one_hom`. -/ @[to_additive "`pure` as a `zero_hom`."] def pure_one_hom : one_hom α (filter α) := ⟨pure, pure_one⟩ @@ -119,7 +122,7 @@ instance : has_inv (filter α) := ⟨map has_inv.inv⟩ end has_inv section has_involutive_inv -variables [has_involutive_inv α] {f : filter α} {s : set α} +variables [has_involutive_inv α] {f g : filter α} {s : set α} @[to_additive] lemma inv_mem_inv (hs : s ∈ f) : s⁻¹ ∈ f⁻¹ := by rwa [mem_inv, inv_preimage, inv_inv] @@ -129,6 +132,17 @@ protected 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 } +localized "attribute [instance] filter.has_involutive_inv filter.has_involutive_neg" in pointwise + +@[simp, to_additive] protected lemma inv_le_inv_iff : f⁻¹ ≤ g⁻¹ ↔ f ≤ g := +⟨λ h, inv_inv f ▸ inv_inv g ▸ filter.inv_le_inv h, filter.inv_le_inv⟩ + +@[to_additive] lemma inv_le_iff_le_inv : f⁻¹ ≤ g ↔ f ≤ g⁻¹ := +by rw [← filter.inv_le_inv_iff, inv_inv] + +@[simp, to_additive] lemma inv_le_self : f⁻¹ ≤ f ↔ f⁻¹ = f := +⟨λ h, h.antisymm $ inv_le_iff_le_inv.1 h, eq.le⟩ + end has_involutive_inv /-! ### Filter addition/multiplication -/ From 3879543c32654fe81bb16486b9837ae3216ede1e Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Wed, 11 Jan 2023 01:12:20 +0000 Subject: [PATCH 0203/1291] =?UTF-8?q?feat(topology/algebra/monoid):=20add?= =?UTF-8?q?=20lemmas=20about=20`=F0=9D=93=9D=20a=20*=20=F0=9D=93=9D=20b`?= =?UTF-8?q?=20(#18099)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit * Add `le_nhds_mul`, `nhds_one_mul_nhds`, `nhds_mul_nhds_one`, `map_mul_right_nhds`, and `map_mul_right_nhds_one`. * Golf `nhds_mul` and generalize it to noncommutative groups. --- src/topology/algebra/group/basic.lean | 36 +++++++++------------------ src/topology/algebra/monoid.lean | 15 ++++++++++- 2 files changed, 26 insertions(+), 25 deletions(-) diff --git a/src/topology/algebra/group/basic.lean b/src/topology/algebra/group/basic.lean index f37ec1565a663..7463aa2e243c2 100644 --- a/src/topology/algebra/group/basic.lean +++ b/src/topology/algebra/group/basic.lean @@ -6,7 +6,6 @@ Authors: Johannes Hölzl, Mario Carneiro, Patrick Massot import group_theory.group_action.conj_act import group_theory.group_action.quotient import group_theory.quotient_group -import order.filter.pointwise import topology.algebra.monoid import topology.algebra.constructions @@ -623,6 +622,11 @@ lemma nhds_translation_mul_inv (x : G) : comap (λ y : G, y * x⁻¹) (𝓝 1) = @[to_additive] lemma map_mul_left_nhds_one (x : G) : map ((*) x) (𝓝 1) = 𝓝 x := by simp +@[simp, to_additive] lemma map_mul_right_nhds (x y : G) : map (λ z, z * x) (𝓝 y) = 𝓝 (y * x) := +(homeomorph.mul_right x).map_nhds_eq y + +@[to_additive] lemma map_mul_right_nhds_one (x : G) : map (λ y, y * x) (𝓝 1) = 𝓝 x := by simp + @[to_additive] lemma filter.has_basis.nhds_of_one {ι : Sort*} {p : ι → Prop} {s : ι → set G} (hb : has_basis (𝓝 1 : filter G) p s) (x : G) : has_basis (𝓝 x) p (λ i, {y | y / x ∈ s i}) := begin @@ -1283,30 +1287,14 @@ end end section -variables [topological_space G] [comm_group G] [topological_group G] +variables [topological_space G] [group G] [topological_group G] -@[to_additive] -lemma nhds_mul (x y : G) : 𝓝 (x * y) = 𝓝 x * 𝓝 y := -filter_eq $ set.ext $ assume s, -begin - rw [← nhds_translation_mul_inv x, ← nhds_translation_mul_inv y, ← nhds_translation_mul_inv (x*y)], - split, - { rintros ⟨t, ht, ts⟩, - rcases exists_nhds_one_split ht with ⟨V, V1, h⟩, - refine ⟨(λa, a * x⁻¹) ⁻¹' V, (λa, a * y⁻¹) ⁻¹' V, - ⟨V, V1, subset.refl _⟩, ⟨V, V1, subset.refl _⟩, _⟩, - rintros a ⟨v, w, v_mem, w_mem, rfl⟩, - apply ts, - simpa [mul_comm, mul_assoc, mul_left_comm] using h (v * x⁻¹) v_mem (w * y⁻¹) w_mem }, - { rintros ⟨a, c, ⟨b, hb, ba⟩, ⟨d, hd, dc⟩, ac⟩, - refine ⟨b ∩ d, inter_mem hb hd, assume v, _⟩, - simp only [preimage_subset_iff, mul_inv_rev, mem_preimage] at *, - rintros ⟨vb, vd⟩, - refine ac ⟨v * y⁻¹, y, _, _, _⟩, - { rw ← mul_assoc _ _ _ at vb, exact ba _ vb }, - { apply dc y, rw mul_right_inv, exact mem_of_mem_nhds hd }, - { simp only [inv_mul_cancel_right] } } -end +@[to_additive] lemma nhds_mul (x y : G) : 𝓝 (x * y) = 𝓝 x * 𝓝 y := +calc 𝓝 (x * y) = map ((*) x) (map (λ a, a * y) (𝓝 1 * 𝓝 1)) : by simp +... = map₂ (λ a b, x * (a * b * y)) (𝓝 1) (𝓝 1) : by rw [← map₂_mul, map_map₂, map_map₂] +... = map₂ (λ a b, x * a * (b * y)) (𝓝 1) (𝓝 1) : by simp only [mul_assoc] +... = 𝓝 x * 𝓝 y : by rw [← map_mul_left_nhds_one x, ← map_mul_right_nhds_one y, ← map₂_mul, + map₂_map_left, map₂_map_right] /-- On a topological group, `𝓝 : G → filter G` can be promoted to a `mul_hom`. -/ @[to_additive "On an additive topological group, `𝓝 : G → filter G` can be promoted to an diff --git a/src/topology/algebra/monoid.lean b/src/topology/algebra/monoid.lean index 53a4e631d7a27..1d4acc363497c 100644 --- a/src/topology/algebra/monoid.lean +++ b/src/topology/algebra/monoid.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro -/ import algebra.big_operators.finprod -import data.set.pointwise.basic +import order.filter.pointwise import topology.algebra.mul_action import algebra.big_operators.pi @@ -101,6 +101,19 @@ lemma filter.tendsto.mul_const (b : M) {c : M} {f : α → M} {l : filter α} (h : tendsto (λ (k:α), f k) l (𝓝 c)) : tendsto (λ (k:α), f k * b) l (𝓝 (c * b)) := h.mul tendsto_const_nhds +@[to_additive] lemma le_nhds_mul (a b : M) : 𝓝 a * 𝓝 b ≤ 𝓝 (a * b) := +by { rw [← map₂_mul, ← map_uncurry_prod, ← nhds_prod_eq], exact continuous_mul.tendsto _ } + +@[simp, to_additive] lemma nhds_one_mul_nhds {M} [mul_one_class M] [topological_space M] + [has_continuous_mul M] (a : M) : 𝓝 (1 : M) * 𝓝 a = 𝓝 a := +((le_nhds_mul _ _).trans_eq $ congr_arg _ (one_mul a)).antisymm $ + le_mul_of_one_le_left' $ pure_le_nhds 1 + +@[simp, to_additive] lemma nhds_mul_nhds_one {M} [mul_one_class M] [topological_space M] + [has_continuous_mul M] (a : M) : 𝓝 a * 𝓝 1 = 𝓝 a := +((le_nhds_mul _ _).trans_eq $ congr_arg _ (mul_one a)).antisymm $ + le_mul_of_one_le_right' $ pure_le_nhds 1 + section tendsto_nhds variables {𝕜 : Type*} From cf4c49c445991489058260d75dae0ff2b1abca28 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 11 Jan 2023 01:12:21 +0000 Subject: [PATCH 0204/1291] chore(data/real/cau_seq_completion): remove use of `parameters` (#18122) This helps with porting --- src/data/real/basic.lean | 6 +- src/data/real/cau_seq_completion.lean | 103 +++++++++++++------------- 2 files changed, 54 insertions(+), 55 deletions(-) diff --git a/src/data/real/basic.lean b/src/data/real/basic.lean index e6365462ab871..a4e18f0825e26 100644 --- a/src/data/real/basic.lean +++ b/src/data/real/basic.lean @@ -25,7 +25,7 @@ open_locale pointwise /-- The type `ℝ` of real numbers constructed as equivalence classes of Cauchy sequences of rational numbers. -/ structure real := of_cauchy :: -(cauchy : @cau_seq.completion.Cauchy ℚ _ _ _ abs _) +(cauchy : cau_seq.completion.Cauchy (abs : ℚ → ℚ)) notation `ℝ` := real attribute [pp_using_anonymous_constructor] real @@ -50,7 +50,7 @@ lemma ext_cauchy {x y : real} : x.cauchy = y.cauchy → x = y := ext_cauchy_iff.2 /-- The real numbers are isomorphic to the quotient of Cauchy sequences on the rationals. -/ -def equiv_Cauchy : ℝ ≃ cau_seq.completion.Cauchy := +def equiv_Cauchy : ℝ ≃ cau_seq.completion.Cauchy abs := ⟨real.cauchy, real.of_cauchy, λ ⟨_⟩, rfl, λ _, rfl⟩ -- irreducible doesn't work for instances: https://github.com/leanprover-community/lean/issues/511 @@ -88,7 +88,7 @@ lemma cauchy_inv : ∀ f, (f⁻¹ : ℝ).cauchy = f.cauchy⁻¹ /-- `real.equiv_Cauchy` as a ring equivalence. -/ @[simps] -def ring_equiv_Cauchy : ℝ ≃+* cau_seq.completion.Cauchy := +def ring_equiv_Cauchy : ℝ ≃+* cau_seq.completion.Cauchy abs := { to_fun := cauchy, inv_fun := of_cauchy, map_add' := cauchy_add, diff --git a/src/data/real/cau_seq_completion.lean b/src/data/real/cau_seq_completion.lean index 40579d05b841d..0769af1d4a5a5 100644 --- a/src/data/real/cau_seq_completion.lean +++ b/src/data/real/cau_seq_completion.lean @@ -16,85 +16,87 @@ namespace cau_seq.completion open cau_seq section -parameters {α : Type*} [linear_ordered_field α] -parameters {β : Type*} [ring β] {abv : β → α} [is_absolute_value abv] +variables {α : Type*} [linear_ordered_field α] +variables {β : Type*} [ring β] (abv : β → α) [is_absolute_value abv] /-- The Cauchy completion of a ring with absolute value. -/ def Cauchy := @quotient (cau_seq _ abv) cau_seq.equiv +variables {abv} + /-- The map from Cauchy sequences into the Cauchy completion. -/ -def mk : cau_seq _ abv → Cauchy := quotient.mk +def mk : cau_seq _ abv → Cauchy abv := quotient.mk -@[simp] theorem mk_eq_mk (f) : @eq Cauchy ⟦f⟧ (mk f) := rfl +@[simp] theorem mk_eq_mk (f) : @eq (Cauchy abv) ⟦f⟧ (mk f) := rfl -theorem mk_eq {f g} : mk f = mk g ↔ f ≈ g := quotient.eq +theorem mk_eq {f g : cau_seq _ abv} : mk f = mk g ↔ f ≈ g := quotient.eq /-- The map from the original ring into the Cauchy completion. -/ -def of_rat (x : β) : Cauchy := mk (const abv x) +def of_rat (x : β) : Cauchy abv := mk (const abv x) -instance : has_zero Cauchy := ⟨of_rat 0⟩ -instance : has_one Cauchy := ⟨of_rat 1⟩ -instance : inhabited Cauchy := ⟨0⟩ +instance : has_zero (Cauchy abv) := ⟨of_rat 0⟩ +instance : has_one (Cauchy abv) := ⟨of_rat 1⟩ +instance : inhabited (Cauchy abv) := ⟨0⟩ -theorem of_rat_zero : of_rat 0 = 0 := rfl -theorem of_rat_one : of_rat 1 = 1 := rfl +theorem of_rat_zero : (of_rat 0 : Cauchy abv) = 0 := rfl +theorem of_rat_one : (of_rat 1 : Cauchy abv) = 1 := rfl -@[simp] theorem mk_eq_zero {f} : mk f = 0 ↔ lim_zero f := +@[simp] theorem mk_eq_zero {f : cau_seq _ abv} : mk f = 0 ↔ lim_zero f := by have : mk f = 0 ↔ lim_zero (f - 0) := quotient.eq; rwa sub_zero at this -instance : has_add Cauchy := +instance : has_add (Cauchy abv) := ⟨quotient.map₂ (+) $ λ f₁ g₁ hf f₂ g₂ hg, add_equiv_add hf hg⟩ @[simp] theorem mk_add (f g : cau_seq β abv) : mk f + mk g = mk (f + g) := rfl -instance : has_neg Cauchy := +instance : has_neg (Cauchy abv) := ⟨quotient.map has_neg.neg $ λ f₁ f₂ hf, neg_equiv_neg hf⟩ @[simp] theorem mk_neg (f : cau_seq β abv) : -mk f = mk (-f) := rfl -instance : has_mul Cauchy := +instance : has_mul (Cauchy abv) := ⟨quotient.map₂ (*) $ λ f₁ g₁ hf f₂ g₂ hg, mul_equiv_mul hf hg⟩ @[simp] theorem mk_mul (f g : cau_seq β abv) : mk f * mk g = mk (f * g) := rfl -instance : has_sub Cauchy := +instance : has_sub (Cauchy abv) := ⟨quotient.map₂ has_sub.sub $ λ f₁ g₁ hf f₂ g₂ hg, sub_equiv_sub hf hg⟩ @[simp] theorem mk_sub (f g : cau_seq β abv) : mk f - mk g = mk (f - g) := rfl -instance {γ : Type*} [has_smul γ β] [is_scalar_tower γ β β] : has_smul γ Cauchy := +instance {γ : Type*} [has_smul γ β] [is_scalar_tower γ β β] : has_smul γ (Cauchy abv) := ⟨λ c, quotient.map ((•) c) $ λ f₁ g₁ hf, smul_equiv_smul _ hf⟩ @[simp] theorem mk_smul {γ : Type*} [has_smul γ β] [is_scalar_tower γ β β] (c : γ) (f : cau_seq β abv) : c • mk f = mk (c • f) := rfl -instance : has_pow Cauchy ℕ := +instance : has_pow (Cauchy abv) ℕ := ⟨λ x n, quotient.map (^ n) (λ f₁ g₁ hf, pow_equiv_pow hf _) x⟩ @[simp] theorem mk_pow (n : ℕ) (f : cau_seq β abv) : mk f ^ n = mk (f ^ n) := rfl -instance : has_nat_cast Cauchy := ⟨λ n, mk n⟩ -instance : has_int_cast Cauchy := ⟨λ n, mk n⟩ +instance : has_nat_cast (Cauchy abv) := ⟨λ n, mk n⟩ +instance : has_int_cast (Cauchy abv) := ⟨λ n, mk n⟩ -@[simp] theorem of_rat_nat_cast (n : ℕ) : of_rat n = n := rfl -@[simp] theorem of_rat_int_cast (z : ℤ) : of_rat z = z := rfl +@[simp] theorem of_rat_nat_cast (n : ℕ) : (of_rat n : Cauchy abv) = n := rfl +@[simp] theorem of_rat_int_cast (z : ℤ) : (of_rat z : Cauchy abv) = z := rfl -theorem of_rat_add (x y : β) : of_rat (x + y) = of_rat x + of_rat y := +theorem of_rat_add (x y : β) : of_rat (x + y) = (of_rat x + of_rat y : Cauchy abv) := congr_arg mk (const_add _ _) -theorem of_rat_neg (x : β) : of_rat (-x) = -of_rat x := +theorem of_rat_neg (x : β) : of_rat (-x) = (-of_rat x : Cauchy abv) := congr_arg mk (const_neg _) -theorem of_rat_mul (x y : β) : of_rat (x * y) = of_rat x * of_rat y := +theorem of_rat_mul (x y : β) : of_rat (x * y) = (of_rat x * of_rat y : Cauchy abv) := congr_arg mk (const_mul _ _) -private lemma zero_def : 0 = mk 0 := rfl +private lemma zero_def : 0 = (mk 0 : Cauchy abv) := rfl -private lemma one_def : 1 = mk 1 := rfl +private lemma one_def : 1 = (mk 1 : Cauchy abv) := rfl -instance : ring Cauchy := +instance : ring (Cauchy abv) := function.surjective.ring mk (surjective_quotient_mk _) zero_def.symm one_def.symm (λ _ _, (mk_add _ _).symm) (λ _ _, (mk_mul _ _).symm) (λ _, (mk_neg _).symm) (λ _ _, (mk_sub _ _).symm) @@ -103,24 +105,23 @@ function.surjective.ring mk (surjective_quotient_mk _) /-- `cau_seq.completion.of_rat` as a `ring_hom` -/ @[simps] -def of_rat_ring_hom : β →+* Cauchy := +def of_rat_ring_hom : β →+* Cauchy abv := { to_fun := of_rat, map_zero' := of_rat_zero, map_one' := of_rat_one, map_add' := of_rat_add, map_mul' := of_rat_mul, } -theorem of_rat_sub (x y : β) : of_rat (x - y) = of_rat x - of_rat y := +theorem of_rat_sub (x y : β) : of_rat (x - y) = (of_rat x - of_rat y : Cauchy abv) := congr_arg mk (const_sub _ _) end section -parameters {α : Type*} [linear_ordered_field α] -parameters {β : Type*} [comm_ring β] {abv : β → α} [is_absolute_value abv] -local notation `Cauchy` := @Cauchy _ _ _ _ abv _ +variables {α : Type*} [linear_ordered_field α] +variables {β : Type*} [comm_ring β] {abv : β → α} [is_absolute_value abv] -instance : comm_ring Cauchy := +instance : comm_ring (Cauchy abv) := function.surjective.comm_ring mk (surjective_quotient_mk _) zero_def.symm one_def.symm (λ _ _, (mk_add _ _).symm) (λ _ _, (mk_mul _ _).symm) (λ _, (mk_neg _).symm) (λ _ _, (mk_sub _ _).symm) @@ -132,15 +133,14 @@ end open_locale classical section -parameters {α : Type*} [linear_ordered_field α] -parameters {β : Type*} [division_ring β] {abv : β → α} [is_absolute_value abv] -local notation `Cauchy` := @Cauchy _ _ _ _ abv _ +variables {α : Type*} [linear_ordered_field α] +variables {β : Type*} [division_ring β] {abv : β → α} [is_absolute_value abv] -instance : has_rat_cast Cauchy := ⟨λ q, of_rat q⟩ +instance : has_rat_cast (Cauchy abv) := ⟨λ q, of_rat q⟩ -@[simp] theorem of_rat_rat_cast (q : ℚ) : of_rat (↑q : β) = (q : Cauchy) := rfl +@[simp] theorem of_rat_rat_cast (q : ℚ) : of_rat (↑q : β) = (q : Cauchy abv) := rfl -noncomputable instance : has_inv Cauchy := +noncomputable instance : has_inv (Cauchy abv) := ⟨λ x, quotient.lift_on x (λ f, mk $ if h : lim_zero f then 0 else inv f h) $ λ f g fg, begin @@ -156,7 +156,7 @@ noncomputable instance : has_inv Cauchy := mul_assoc, Ig', mul_one] } end⟩ -@[simp] theorem inv_zero : (0 : Cauchy)⁻¹ = 0 := +@[simp] theorem inv_zero : (0 : Cauchy abv)⁻¹ = 0 := congr_arg mk $ by rw dif_pos; [refl, exact zero_lim_zero] @[simp] theorem inv_mk {f} (hf) : (@mk α _ β _ abv _ f)⁻¹ = mk (inv f hf) := @@ -167,26 +167,26 @@ have lim_zero (1 - 0), from setoid.symm h, have lim_zero 1, by simpa, one_ne_zero $ const_lim_zero.1 this -lemma zero_ne_one : (0 : Cauchy) ≠ 1 := +lemma zero_ne_one : (0 : Cauchy abv) ≠ 1 := λ h, cau_seq_zero_ne_one $ mk_eq.1 h -protected theorem inv_mul_cancel {x : Cauchy} : x ≠ 0 → x⁻¹ * x = 1 := +protected theorem inv_mul_cancel {x : Cauchy abv} : x ≠ 0 → x⁻¹ * x = 1 := quotient.induction_on x $ λ f hf, begin simp at hf, simp [hf], exact quotient.sound (cau_seq.inv_mul_cancel hf) end -protected theorem mul_inv_cancel {x : Cauchy} : x ≠ 0 → x * x⁻¹ = 1 := +protected theorem mul_inv_cancel {x : Cauchy abv} : x ≠ 0 → x * x⁻¹ = 1 := quotient.induction_on x $ λ f hf, begin simp at hf, simp [hf], exact quotient.sound (cau_seq.mul_inv_cancel hf) end -theorem of_rat_inv (x : β) : of_rat (x⁻¹) = ((of_rat x)⁻¹ : Cauchy) := +theorem of_rat_inv (x : β) : of_rat (x⁻¹) = ((of_rat x)⁻¹ : Cauchy abv) := congr_arg mk $ by split_ifs with h; [simp [const_lim_zero.1 h], refl] /-- The Cauchy completion forms a division ring. -/ -noncomputable instance : division_ring Cauchy := +noncomputable instance : division_ring (Cauchy abv) := { inv := has_inv.inv, mul_inv_cancel := λ x, cau_seq.completion.mul_inv_cancel, exists_pair_ne := ⟨0, 1, zero_ne_one⟩, @@ -196,7 +196,7 @@ noncomputable instance : division_ring Cauchy := by rw [rat.cast_mk', of_rat_mul, of_rat_int_cast, of_rat_inv, of_rat_nat_cast], .. Cauchy.ring } -theorem of_rat_div (x y : β) : of_rat (x / y) = (of_rat x / of_rat y : Cauchy) := +theorem of_rat_div (x y : β) : of_rat (x / y) = (of_rat x / of_rat y : Cauchy abv) := by simp only [div_eq_mul_inv, of_rat_inv, of_rat_mul] /-- Show the first 10 items of a representative of this equivalence class of cauchy sequences. @@ -204,7 +204,7 @@ by simp only [div_eq_mul_inv, of_rat_inv, of_rat_mul] The representative chosen is the one passed in the VM to `quot.mk`, so two cauchy sequences converging to the same number may be printed differently. -/ -meta instance [has_repr β] : has_repr Cauchy := +meta instance [has_repr β] : has_repr (Cauchy abv) := { repr := λ r, let N := 10, seq := r.unquot in "(sorry /- " ++ (", ".intercalate $ (list.range N).map $ repr ∘ seq) ++ ", ... -/)" } @@ -212,12 +212,11 @@ meta instance [has_repr β] : has_repr Cauchy := end section -parameters {α : Type*} [linear_ordered_field α] -parameters {β : Type*} [field β] {abv : β → α} [is_absolute_value abv] -local notation `Cauchy` := @Cauchy _ _ _ _ abv _ +variables {α : Type*} [linear_ordered_field α] +variables {β : Type*} [field β] {abv : β → α} [is_absolute_value abv] /-- The Cauchy completion forms a field. -/ -noncomputable instance : field Cauchy := +noncomputable instance : field (Cauchy abv) := { .. Cauchy.division_ring, .. Cauchy.comm_ring } From ccad6d5093bd2f5c6ca621fc74674cce51355af6 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 11 Jan 2023 04:21:24 +0000 Subject: [PATCH 0205/1291] chore(data/real/basic): move some data out of real.comm_ring (#18118) This pulls out the `int_cast`, `nat_cast`, and `has_sub` data typeclasses from `real.comm_ring`. The `of_cauchy_sub` and `cauchy_sub` lemmas are also new. This is intended to be the safe half of #8146, and is motivated by shrinking the diff there to bisect the issue. Co-authored-by: Scott Morrison --- src/data/real/basic.lean | 53 ++++++++++++++++++++++------------------ 1 file changed, 29 insertions(+), 24 deletions(-) diff --git a/src/data/real/basic.lean b/src/data/real/basic.lean index a4e18f0825e26..8117c41595b3b 100644 --- a/src/data/real/basic.lean +++ b/src/data/real/basic.lean @@ -66,12 +66,15 @@ instance : has_one ℝ := ⟨one⟩ instance : has_add ℝ := ⟨add⟩ instance : has_neg ℝ := ⟨neg⟩ instance : has_mul ℝ := ⟨mul⟩ +instance : has_sub ℝ := ⟨λ a b, a + (-b)⟩ noncomputable instance : has_inv ℝ := ⟨inv'⟩ lemma of_cauchy_zero : (⟨0⟩ : ℝ) = 0 := show _ = zero, by rw zero lemma of_cauchy_one : (⟨1⟩ : ℝ) = 1 := show _ = one, by rw one lemma of_cauchy_add (a b) : (⟨a + b⟩ : ℝ) = ⟨a⟩ + ⟨b⟩ := show _ = add _ _, by rw add lemma of_cauchy_neg (a) : (⟨-a⟩ : ℝ) = -⟨a⟩ := show _ = neg _, by rw neg +lemma of_cauchy_sub (a b) : (⟨a - b⟩ : ℝ) = ⟨a⟩ - ⟨b⟩ := +by { rw [sub_eq_add_neg, of_cauchy_add, of_cauchy_neg], refl } lemma of_cauchy_mul (a b) : (⟨a * b⟩ : ℝ) = ⟨a⟩ * ⟨b⟩ := show _ = mul _ _, by rw mul lemma of_cauchy_inv {f} : (⟨f⁻¹⟩ : ℝ) = ⟨f⟩⁻¹ := show _ = inv' _, by rw inv' @@ -83,17 +86,22 @@ lemma cauchy_neg : ∀ a, (-a : ℝ).cauchy = -a.cauchy | ⟨a⟩ := show (neg _).cauchy = _, by rw neg lemma cauchy_mul : ∀ a b, (a * b : ℝ).cauchy = a.cauchy * b.cauchy | ⟨a⟩ ⟨b⟩ := show (mul _ _).cauchy = _, by rw mul +lemma cauchy_sub : ∀ a b, (a - b : ℝ).cauchy = a.cauchy - b.cauchy +| ⟨a⟩ ⟨b⟩ := by { rw [sub_eq_add_neg, ←cauchy_neg, ←cauchy_add], refl } lemma cauchy_inv : ∀ f, (f⁻¹ : ℝ).cauchy = f.cauchy⁻¹ | ⟨f⟩ := show (inv' _).cauchy = _, by rw inv' -/-- `real.equiv_Cauchy` as a ring equivalence. -/ -@[simps] -def ring_equiv_Cauchy : ℝ ≃+* cau_seq.completion.Cauchy abs := -{ to_fun := cauchy, - inv_fun := of_cauchy, - map_add' := cauchy_add, - map_mul' := cauchy_mul, - ..equiv_Cauchy } +instance : has_nat_cast ℝ := { nat_cast := λ n, ⟨n⟩ } +instance : has_int_cast ℝ := { int_cast := λ z, ⟨z⟩ } +instance : has_rat_cast ℝ := { rat_cast := λ q, ⟨q⟩ } + +lemma of_cauchy_nat_cast (n : ℕ) : (⟨n⟩ : ℝ) = n := rfl +lemma of_cauchy_int_cast (z : ℤ) : (⟨z⟩ : ℝ) = z := rfl +lemma of_cauchy_rat_cast (q : ℚ) : (⟨q⟩ : ℝ) = q := rfl + +lemma cauchy_nat_cast (n : ℕ) : (n : ℝ).cauchy = n := rfl +lemma cauchy_int_cast (z : ℤ) : (z : ℝ).cauchy = z := rfl +lemma cauchy_rat_cast (q : ℚ) : (q : ℝ).cauchy = q := rfl instance : comm_ring ℝ := begin @@ -102,30 +110,28 @@ begin mul := (*), add := (+), neg := @has_neg.neg ℝ _, - sub := λ a b, a + (-b), - nat_cast := λ n, ⟨n⟩, - int_cast := λ n, ⟨n⟩, + sub := @has_sub.sub ℝ _, npow := @npow_rec ℝ ⟨1⟩ ⟨(*)⟩, nsmul := @nsmul_rec ℝ ⟨0⟩ ⟨(+)⟩, - zsmul := @zsmul_rec ℝ ⟨0⟩ ⟨(+)⟩ ⟨@has_neg.neg ℝ _⟩ }; + zsmul := @zsmul_rec ℝ ⟨0⟩ ⟨(+)⟩ ⟨@has_neg.neg ℝ _⟩, + ..real.has_nat_cast, + ..real.has_int_cast, }; repeat { rintro ⟨_⟩, }; try { refl }; simp [← of_cauchy_zero, ← of_cauchy_one, ←of_cauchy_add, ←of_cauchy_neg, ←of_cauchy_mul, - λ n, show @coe ℕ ℝ ⟨_⟩ n = ⟨n⟩, from rfl]; + λ n, show @coe ℕ ℝ ⟨_⟩ n = ⟨n⟩, from rfl, has_nat_cast.nat_cast, has_int_cast.int_cast]; apply add_assoc <|> apply add_comm <|> apply mul_assoc <|> apply mul_comm <|> apply left_distrib <|> apply right_distrib <|> apply sub_eq_add_neg <|> skip, end - -instance : has_rat_cast ℝ := { rat_cast := λ q, ⟨q⟩ } - -lemma of_cauchy_nat_cast (n : ℕ) : (⟨n⟩ : ℝ) = n := rfl -lemma of_cauchy_int_cast (z : ℤ) : (⟨z⟩ : ℝ) = z := rfl -lemma of_cauchy_rat_cast (q : ℚ) : (⟨q⟩ : ℝ) = q := rfl - -lemma cauchy_nat_cast (n : ℕ) : (n : ℝ).cauchy = n := rfl -lemma cauchy_int_cast (z : ℤ) : (z : ℝ).cauchy = z := rfl -lemma cauchy_rat_cast (q : ℚ) : (q : ℝ).cauchy = q := rfl +/-- `real.equiv_Cauchy` as a ring equivalence. -/ +@[simps] +def ring_equiv_Cauchy : ℝ ≃+* cau_seq.completion.Cauchy abs := +{ to_fun := cauchy, + inv_fun := of_cauchy, + map_add' := cauchy_add, + map_mul' := cauchy_mul, + ..equiv_Cauchy } /-! Extra instances to short-circuit type class resolution. @@ -149,7 +155,6 @@ instance : comm_monoid ℝ := by apply_instance instance : monoid ℝ := by apply_instance instance : comm_semigroup ℝ := by apply_instance instance : semigroup ℝ := by apply_instance -instance : has_sub ℝ := by apply_instance instance : inhabited ℝ := ⟨0⟩ /-- The real numbers are a `*`-ring, with the trivial `*`-structure. -/ From 6a033cb3d188a12ca5c509b33e2eaac1c61916cd Mon Sep 17 00:00:00 2001 From: loefflerd Date: Wed, 11 Jan 2023 07:28:27 +0000 Subject: [PATCH 0206/1291] feat(number_theory/bernoulli): Fourier coefficients of Bernoulli polynomials (#17990) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Add the integral formula `∫ x in 0..1, exp(-2 * π * I * n * x) * B k x = -k! / (2 * π * I * n) ^ k`, where `B k` is the `k`-th Bernoulli polynomial. Co-authored-by: Junyan Xu --- src/analysis/fourier/add_circle.lean | 123 ++++++++++++++++++- src/data/real/basic.lean | 2 + src/geometry/manifold/instances/real.lean | 4 +- src/measure_theory/integral/periodic.lean | 3 +- src/number_theory/well_approximable.lean | 2 +- src/number_theory/zeta_values.lean | 136 ++++++++++++++++++++++ src/topology/instances/add_circle.lean | 3 +- 7 files changed, 261 insertions(+), 12 deletions(-) create mode 100644 src/number_theory/zeta_values.lean diff --git a/src/analysis/fourier/add_circle.lean b/src/analysis/fourier/add_circle.lean index 0a24bd245014b..fa8ed0c38f94b 100644 --- a/src/analysis/fourier/add_circle.lean +++ b/src/analysis/fourier/add_circle.lean @@ -28,8 +28,13 @@ This file contains basic results on Fourier series for functions on the additive from `add_circle T` to `ℂ`. * `fourier_basis` is the Hilbert basis of `Lp ℂ 2 haar_add_circle` given by the images of the monomials `fourier n`. -* `fourier_coeff f n`, for `f : add_circle T → ℂ`, is the `n`-th Fourier coefficient of `f` - (defined as an integral over `add_circle T`). +* `fourier_coeff f n`, for `f : add_circle T → E` (with `E` a complete normed `ℂ`-vector space), is + the `n`-th Fourier coefficient of `f`, defined as an integral over `add_circle T`. The lemma + `fourier_coeff_eq_interval_integral` expresses this as an integral over `[a, a + T]` for any real + `a`. +* `fourier_coeff_on`, for `f : ℝ → E` and `a < b` reals, is the `n`-th Fourier + coefficient of the unique periodic function of period `b - a` which agrees with `f` on `(a, b]`. + The lemma `fourier_coeff_on_eq_integral` expresses this as an integral over `[a, b]`. ## Main statements @@ -299,8 +304,8 @@ variables {E : Type} [normed_add_comm_group E] [normed_space ℂ E] [complete_sp def fourier_coeff (f : add_circle T → E) (n : ℤ) : E := ∫ (t : add_circle T), fourier (-n) t • f t ∂ haar_add_circle -/-- The Fourier coefficients of a function can be computed as an integral -over `[a, a + T]` for any real `a`. -/ +/-- The Fourier coefficients of a function on `add_circle T` can be computed as an integral +over `[a, a + T]`, for any real `a`. -/ lemma fourier_coeff_eq_interval_integral (f : add_circle T → E) (n : ℤ) (a : ℝ) : fourier_coeff f n = (1 / T) • ∫ x in a .. a + T, @fourier T (-n) x • f x := begin @@ -312,6 +317,47 @@ begin ←smul_assoc, smul_eq_mul, one_div_mul_cancel hT.out.ne', one_smul], end +lemma fourier_coeff.const_smul (f : add_circle T → E) (c : ℂ) (n : ℤ) : + fourier_coeff (c • f) n = c • fourier_coeff f n := +by simp_rw [fourier_coeff, pi.smul_apply, ←smul_assoc, smul_eq_mul, mul_comm, ←smul_eq_mul, + smul_assoc, integral_smul] + +lemma fourier_coeff.const_mul (f : add_circle T → ℂ) (c : ℂ) (n : ℤ) : + fourier_coeff (λ x, c * f x) n = c * fourier_coeff f n := +fourier_coeff.const_smul f c n + +omit hT + +/-- For a function on `ℝ`, the Fourier coefficients of `f` on `[a, b]` are defined as the +Fourier coefficients of the unique periodic function agreeing with `f` on `Ioc a b`. -/ +def fourier_coeff_on {a b : ℝ} (hab : a < b) (f : ℝ → E) (n : ℤ) : E := +begin + haveI := fact.mk (by linarith : 0 < b - a), + exact fourier_coeff (add_circle.lift_Ioc (b - a) a f) n +end + +lemma fourier_coeff_on_eq_integral {a b : ℝ} (f : ℝ → E) (n : ℤ) (hab : a < b) : + fourier_coeff_on hab f n = + (1 / (b - a)) • ∫ x in a ..b, fourier (-n) (x : add_circle (b - a)) • f x := +begin + rw [fourier_coeff_on, fourier_coeff_eq_interval_integral _ _ a], + congr' 1, + rw [add_sub, add_sub_cancel'], + simp_rw interval_integral.integral_of_le hab.le, + refine set_integral_congr measurable_set_Ioc (λ x hx, _), + dsimp only, + rwa [lift_Ioc_coe_apply], + rwa [add_sub, add_sub_cancel'], +end + +lemma fourier_coeff_on.const_smul {a b : ℝ} (f : ℝ → E) (c : ℂ) (n : ℤ) (hab : a < b) : + fourier_coeff_on hab (c • f) n = c • fourier_coeff_on hab f n := +by apply fourier_coeff.const_smul + +lemma fourier_coeff_on.const_mul {a b : ℝ} (f : ℝ → ℂ) (c : ℂ) (n : ℤ) (hab : a < b) : + fourier_coeff_on hab (λ x, c * f x) n = c * fourier_coeff_on hab f n := +fourier_coeff_on.const_smul _ _ _ _ + end fourier_coeff section fourier_L2 @@ -397,3 +443,72 @@ lemma has_pointwise_sum_fourier_series_of_summable end convergence end scope_hT + + +section deriv + +open complex interval_integral +open_locale interval + +variables (T) + +lemma has_deriv_at_fourier (n : ℤ) (x : ℝ) : has_deriv_at (λ y:ℝ, fourier n (y : add_circle T)) + (2 * π * I * n / T * fourier n (x : add_circle T)) x := +begin + simp_rw [fourier_coe_apply], + refine (_ : has_deriv_at (λ y, exp (2 * π * I * n * y / T)) _ _).comp_of_real, + rw (λ α β, by ring : ∀ (α β : ℂ), α * exp β = exp β * α), + refine (has_deriv_at_exp _).comp x _, + convert has_deriv_at_mul_const (2 * ↑π * I * ↑n / T), + ext1 y, ring, +end + +lemma has_deriv_at_fourier_neg (n : ℤ) (x : ℝ) : + has_deriv_at (λ y:ℝ, fourier (-n) (y : add_circle T)) + (-2 * π * I * n / T * fourier (-n) (x : add_circle T)) x := +by simpa using has_deriv_at_fourier T (-n) x + +variables {T} + +lemma has_antideriv_at_fourier_neg (hT : fact (0 < T)) {n : ℤ} (hn : n ≠ 0) (x : ℝ) : + has_deriv_at (λ (y : ℝ), (T : ℂ) / (-2 * π * I * n) * fourier (-n) (y : add_circle T)) + (fourier (-n) (x : add_circle T)) x := +begin + convert (has_deriv_at_fourier_neg T n x).div_const (-2 * π * I * n / T) using 1, + { ext1 y, rw div_div_eq_mul_div, ring, }, + { rw mul_div_cancel_left, + simp only [ne.def, div_eq_zero_iff, neg_eq_zero, mul_eq_zero, bit0_eq_zero, one_ne_zero, + of_real_eq_zero, false_or, int.cast_eq_zero, not_or_distrib], + exact ⟨⟨⟨real.pi_ne_zero, I_ne_zero⟩, hn⟩, hT.out.ne'⟩ }, +end + +/-- Express Fourier coefficients of `f` on an interval in terms of those of its derivative. -/ +lemma fourier_coeff_on_of_has_deriv_at {a b : ℝ} (hab : a < b) {f f' : ℝ → ℂ} {n : ℤ} + (hn : n ≠ 0) (hf : ∀ x, x ∈ [a, b] → has_deriv_at f (f' x) x) + (hf' : interval_integrable f' volume a b) : + fourier_coeff_on hab f n = + 1 / (-2 * π * I * n) * (fourier (-n) (a : add_circle (b - a)) * (f b - f a) + - (b - a) * fourier_coeff_on hab f' n) := +begin + rw ←of_real_sub, + have hT : fact (0 < b - a) := ⟨by linarith⟩, + simp_rw [fourier_coeff_on_eq_integral, smul_eq_mul, real_smul, of_real_div, of_real_one], + conv { for (fourier _ _ * _) [1, 2, 3] { rw mul_comm } }, + rw integral_mul_deriv_eq_deriv_mul hf (λ x hx, has_antideriv_at_fourier_neg hT hn x) hf' + (((map_continuous (fourier (-n))).comp (add_circle.continuous_mk' _)).interval_integrable _ _), + dsimp only, + have : ∀ (u v w : ℂ), u * ( (b - a : ℝ) / v * w) = (b - a : ℝ) / v * (u * w) := by {intros, ring}, + conv in (interval_integral _ _ _ _) { congr, funext, rw this, }, + rw (by ring : ((b - a : ℝ) : ℂ) / ((-2) * π * I * n) + = ((b - a : ℝ) : ℂ) * (1 / ((-2) * π * I * n))), + have s2 : (b : add_circle (b - a)) = (a : add_circle (b - a)), + { simpa using coe_add_period (b - a) a, }, + rw [s2, integral_const_mul, ←sub_mul, mul_sub, mul_sub], + congr' 1, + { conv_lhs {rw [mul_comm, mul_div, mul_one]}, + rw [div_eq_iff (of_real_ne_zero.mpr hT.out.ne')], + ring, }, + { ring, }, +end + +end deriv diff --git a/src/data/real/basic.lean b/src/data/real/basic.lean index 8117c41595b3b..76aa36a70ac09 100644 --- a/src/data/real/basic.lean +++ b/src/data/real/basic.lean @@ -237,6 +237,8 @@ end protected theorem zero_lt_one : (0 : ℝ) < 1 := by convert rat_cast_lt.2 zero_lt_one; simp [←of_cauchy_rat_cast, of_cauchy_one, of_cauchy_zero] +protected lemma fact_zero_lt_one : fact ((0 : ℝ) < 1) := ⟨real.zero_lt_one⟩ + protected theorem mul_pos {a b : ℝ} : 0 < a → 0 < b → 0 < a * b := begin induction a using real.ind_mk with a, diff --git a/src/geometry/manifold/instances/real.lean b/src/geometry/manifold/instances/real.lean index 65f01d24a5f95..533dcf5c7af9a 100644 --- a/src/geometry/manifold/instances/real.lean +++ b/src/geometry/manifold/instances/real.lean @@ -303,9 +303,7 @@ end /-! Register the manifold structure on `Icc 0 1`, and also its zero and one. -/ section -lemma fact_zero_lt_one : fact ((0 : ℝ) < 1) := ⟨zero_lt_one⟩ - -local attribute [instance] fact_zero_lt_one +local attribute [instance] real.fact_zero_lt_one instance : charted_space (euclidean_half_space 1) (Icc (0 : ℝ) 1) := by apply_instance instance : smooth_manifold_with_corners (𝓡∂ 1) (Icc (0 : ℝ) 1) := by apply_instance diff --git a/src/measure_theory/integral/periodic.lean b/src/measure_theory/integral/periodic.lean index 9f84cf6475e36..f1f77629745cc 100644 --- a/src/measure_theory/integral/periodic.lean +++ b/src/measure_theory/integral/periodic.lean @@ -188,8 +188,7 @@ end end add_circle namespace unit_add_circle -private lemma fact_zero_lt_one : fact ((0:ℝ) < 1) := ⟨zero_lt_one⟩ -local attribute [instance] fact_zero_lt_one +local attribute [instance] real.fact_zero_lt_one noncomputable instance measure_space : measure_space unit_add_circle := add_circle.measure_space 1 diff --git a/src/number_theory/well_approximable.lean b/src/number_theory/well_approximable.lean index 69255ac38e2c6..621d66baf8be1 100644 --- a/src/number_theory/well_approximable.lean +++ b/src/number_theory/well_approximable.lean @@ -156,7 +156,7 @@ lemma mem_approx_add_order_of_iff {δ : ℝ} {x : unit_add_circle} {n : ℕ} (hn x ∈ approx_add_order_of unit_add_circle n δ ↔ ∃ m < n, gcd m n = 1 ∧ ‖x - ↑((m : ℝ) / n)‖ < δ := begin - haveI : fact ((0 : ℝ) < 1) := ⟨zero_lt_one⟩, + haveI := real.fact_zero_lt_one, simp only [mem_approx_add_order_of_iff, mem_set_of_eq, ball, exists_prop, dist_eq_norm, add_circle.add_order_of_eq_pos_iff hn, mul_one], split, diff --git a/src/number_theory/zeta_values.lean b/src/number_theory/zeta_values.lean new file mode 100644 index 0000000000000..6484074af7da9 --- /dev/null +++ b/src/number_theory/zeta_values.lean @@ -0,0 +1,136 @@ +/- +Copyright (c) 2022 David Loeffler. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: David Loeffler +-/ + +import number_theory.bernoulli_polynomials +import analysis.special_functions.integrals +import measure_theory.integral.interval_integral +import analysis.fourier.add_circle + +/-! +# Critical values of the Riemann zeta function + +In this file we evaluate the Fourier coefficients of Bernoulli polynomials on the interval `[0, 1]`. +In a future iteration this will be used to evaluate critical values of the Riemann zeta function +(and other Dirichlet L-functions). +-/ + +noncomputable theory +open_locale nat real interval +open complex measure_theory set interval_integral + +section bernoulli_fun_props +/-! Simple properties of the Bernoulli polynomial, as a function `ℝ → ℝ`. -/ + +/-- The function `x ↦ Bₖ(x) : ℝ → ℝ`. -/ +def bernoulli_fun (k : ℕ) (x : ℝ) : ℝ := +(polynomial.map (algebra_map ℚ ℝ) (polynomial.bernoulli k)).eval x + +lemma bernoulli_fun_eval_zero (k : ℕ) : bernoulli_fun k 0 = bernoulli k := +by rw [bernoulli_fun, polynomial.eval_zero_map, polynomial.bernoulli_eval_zero, eq_rat_cast] + +lemma bernoulli_fun_endpoints_eq_of_ne_one {k : ℕ} (hk : k ≠ 1) : + bernoulli_fun k 1 = bernoulli_fun k 0 := +by rw [bernoulli_fun_eval_zero, bernoulli_fun, polynomial.eval_one_map, + polynomial.bernoulli_eval_one, bernoulli_eq_bernoulli'_of_ne_one hk, eq_rat_cast] + +lemma bernoulli_fun_eval_one (k : ℕ) : bernoulli_fun k 1 = bernoulli_fun k 0 + ite (k = 1) 1 0 := +begin + rw [bernoulli_fun, bernoulli_fun_eval_zero, polynomial.eval_one_map, + polynomial.bernoulli_eval_one], + split_ifs, + { rw [h, bernoulli_one, bernoulli'_one, eq_rat_cast], + push_cast, ring }, + { rw [bernoulli_eq_bernoulli'_of_ne_one h, add_zero, eq_rat_cast], } +end + +lemma has_deriv_at_bernoulli_fun (k : ℕ) (x : ℝ) : + has_deriv_at (bernoulli_fun k) (k * bernoulli_fun (k - 1) x) x := +begin + convert ((polynomial.bernoulli k).map $ algebra_map ℚ ℝ).has_deriv_at x using 1, + simp only [bernoulli_fun, polynomial.derivative_map, polynomial.derivative_bernoulli k, + polynomial.map_mul, polynomial.map_nat_cast, polynomial.eval_mul, polynomial.eval_nat_cast], +end + +lemma antideriv_bernoulli_fun (k : ℕ) (x : ℝ) : + has_deriv_at (λ x, (bernoulli_fun (k + 1) x) / (k + 1)) (bernoulli_fun k x) x := +begin + convert (has_deriv_at_bernoulli_fun (k + 1) x).div_const _, + field_simp [nat.cast_add_one_ne_zero k], + ring, +end + +lemma integral_bernoulli_fun_eq_zero {k : ℕ} (hk : k ≠ 0) : + ∫ (x : ℝ) in 0..1, bernoulli_fun k x = 0 := +begin + rw integral_eq_sub_of_has_deriv_at (λ x hx, antideriv_bernoulli_fun k x) + ((polynomial.continuous _).interval_integrable _ _), + dsimp only, + rw bernoulli_fun_eval_one, + split_ifs, + { exfalso, exact hk (nat.succ_inj'.mp h), }, { simp }, +end + +end bernoulli_fun_props + +section bernoulli_fourier_coeffs +/-! Compute the Fourier coefficients of the Bernoulli functions via integration by parts. -/ + +local attribute [instance] real.fact_zero_lt_one + +/-- The `n`-th Fourier coefficient of the `k`-th Bernoulli function on the interval `[0, 1]`. -/ +def bernoulli_fourier_coeff (k : ℕ) (n : ℤ) : ℂ := +fourier_coeff_on zero_lt_one (λ x, bernoulli_fun k x) n + +/-- Recurrence relation (in `k`) for the `n`-th Fourier coefficient of `Bₖ`. -/ +lemma bernoulli_fourier_coeff_recurrence (k : ℕ) {n : ℤ} (hn : n ≠ 0) : + bernoulli_fourier_coeff k n = 1 / ((-2) * π * I * n) * + (ite (k = 1) 1 0 - k * bernoulli_fourier_coeff (k - 1) n) := +begin + unfold bernoulli_fourier_coeff, + rw [fourier_coeff_on_of_has_deriv_at zero_lt_one + hn (λ x hx, (has_deriv_at_bernoulli_fun k x).of_real_comp) + ((continuous_of_real.comp $ continuous_const.mul + $ polynomial.continuous _).interval_integrable _ _)], + dsimp only, + simp_rw [of_real_one, of_real_zero, sub_zero, one_mul], + rw [quotient_add_group.coe_zero, fourier_eval_zero, one_mul, + ←of_real_sub, bernoulli_fun_eval_one, add_sub_cancel'], + congr' 2, + { split_ifs, all_goals { simp only [of_real_one, of_real_zero, one_mul]}, }, + { simp_rw [of_real_mul, of_real_nat_cast, fourier_coeff_on.const_mul] }, +end + +/-- The Fourier coefficients of `B₀(x) = 1`. -/ +lemma bernoulli_zero_fourier_coeff {n : ℤ} (hn : n ≠ 0) : bernoulli_fourier_coeff 0 n = 0 := +by simpa using bernoulli_fourier_coeff_recurrence 0 hn + +/-- The `0`-th Fourier coefficient of `Bₖ(x)`. -/ +lemma bernoulli_fourier_coeff_zero {k : ℕ} (hk : k ≠ 0) : bernoulli_fourier_coeff k 0 = 0 := +by simp_rw [bernoulli_fourier_coeff, fourier_coeff_on_eq_integral, neg_zero, fourier_zero, sub_zero, + div_one, one_smul, interval_integral.integral_of_real, integral_bernoulli_fun_eq_zero hk, + of_real_zero] + +lemma bernoulli_fourier_coeff_eq {k : ℕ} (hk : k ≠ 0) (n : ℤ) : + bernoulli_fourier_coeff k n = - k! / (2 * π * I * n) ^ k := +begin + rcases eq_or_ne n 0 with rfl|hn, + { rw [bernoulli_fourier_coeff_zero hk, int.cast_zero, mul_zero, + zero_pow' _ hk, div_zero] }, + refine nat.le_induction _ (λ k hk h'k, _) k (nat.one_le_iff_ne_zero.mpr hk), + { rw bernoulli_fourier_coeff_recurrence 1 hn, + simp only [nat.cast_one, tsub_self, neg_mul, one_mul, eq_self_iff_true, if_true, + nat.factorial_one, pow_one, inv_I, mul_neg], + rw [bernoulli_zero_fourier_coeff hn, sub_zero, mul_one, div_neg, neg_div], }, + { rw [bernoulli_fourier_coeff_recurrence (k + 1) hn, nat.add_sub_cancel k 1], + split_ifs, + { exfalso, exact (ne_of_gt (nat.lt_succ_iff.mpr hk)) h,}, + { rw [h'k, nat.factorial_succ, zero_sub, nat.cast_mul, pow_add, pow_one, neg_div, + mul_neg, mul_neg, mul_neg, neg_neg, neg_mul, neg_mul, neg_mul, div_neg], + field_simp [int.cast_ne_zero.mpr hn, I_ne_zero], + ring_nf, } } +end + +end bernoulli_fourier_coeffs diff --git a/src/topology/instances/add_circle.lean b/src/topology/instances/add_circle.lean index 7049ede5b7f5e..a67ddcb2eb982 100644 --- a/src/topology/instances/add_circle.lean +++ b/src/topology/instances/add_circle.lean @@ -447,8 +447,7 @@ instance : second_countable_topology (add_circle p) := quotient_add_group.second end add_circle -private lemma fact_zero_lt_one : fact ((0:ℝ) < 1) := ⟨zero_lt_one⟩ -local attribute [instance] fact_zero_lt_one +local attribute [instance] real.fact_zero_lt_one /-- The unit circle `ℝ ⧸ ℤ`. -/ @[derive [compact_space, normal_space, second_countable_topology]] From be24ec5de6701447e5df5ca75400ffee19d65659 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Wed, 11 Jan 2023 07:28:28 +0000 Subject: [PATCH 0207/1291] chore(*): add mathlib4 synchronization comments (#18094) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.bounds` * `algebra.free_monoid.basic` * `algebra.module.hom` * `algebra.module.pi` * `algebra.order.archimedean` * `algebra.order.floor` * `algebra.punit_instances` * `algebra.star.basic` * `algebra.star.pi` * `algebra.star.prod` * `algebra.star.unitary` * `data.list.big_operators.basic` * `data.list.big_operators.lemmas` * `data.list.chain` * `data.list.count` * `data.list.dedup` * `data.list.duplicate` * `data.list.join` * `data.list.lattice` * `data.list.nodup` * `data.list.pairwise` * `data.list.permutation` * `data.list.prod_sigma` * `data.list.range` * `data.list.rdrop` * `data.list.zip` * `data.nat.bitwise` * `data.rat.floor` * `data.real.cau_seq` * `data.set.pointwise.basic` * `logic.equiv.local_equiv` --- src/algebra/bounds.lean | 3 +++ src/algebra/free_monoid/basic.lean | 3 +++ src/algebra/module/hom.lean | 3 +++ src/algebra/module/pi.lean | 3 +++ src/algebra/order/archimedean.lean | 3 +++ src/algebra/order/floor.lean | 3 +++ src/algebra/punit_instances.lean | 3 +++ src/algebra/star/basic.lean | 3 +++ src/algebra/star/pi.lean | 3 +++ src/algebra/star/prod.lean | 3 +++ src/algebra/star/unitary.lean | 3 +++ src/data/list/big_operators/basic.lean | 3 +++ src/data/list/big_operators/lemmas.lean | 5 ++++- src/data/list/chain.lean | 3 +++ src/data/list/count.lean | 3 +++ src/data/list/dedup.lean | 3 +++ src/data/list/duplicate.lean | 3 +++ src/data/list/join.lean | 3 +++ src/data/list/lattice.lean | 3 +++ src/data/list/nodup.lean | 3 +++ src/data/list/pairwise.lean | 3 +++ src/data/list/permutation.lean | 3 +++ src/data/list/prod_sigma.lean | 3 +++ src/data/list/range.lean | 3 +++ src/data/list/rdrop.lean | 3 +++ src/data/list/zip.lean | 3 +++ src/data/nat/bitwise.lean | 3 +++ src/data/rat/floor.lean | 3 +++ src/data/real/cau_seq.lean | 3 +++ src/data/set/pointwise/basic.lean | 3 +++ src/logic/equiv/local_equiv.lean | 3 +++ 31 files changed, 94 insertions(+), 1 deletion(-) diff --git a/src/algebra/bounds.lean b/src/algebra/bounds.lean index 9e672062d77e7..6d0286f9ff436 100644 --- a/src/algebra/bounds.lean +++ b/src/algebra/bounds.lean @@ -12,6 +12,9 @@ import order.conditionally_complete_lattice.basic /-! # Upper/lower bounds in ordered monoids and groups +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove a few facts like “`-s` is bounded above iff `s` is bounded below” (`bdd_above_neg`). -/ diff --git a/src/algebra/free_monoid/basic.lean b/src/algebra/free_monoid/basic.lean index 224f4956db049..7f3a305ab6003 100644 --- a/src/algebra/free_monoid/basic.lean +++ b/src/algebra/free_monoid/basic.lean @@ -8,6 +8,9 @@ import data.list.big_operators.basic /-! # Free monoid over a given alphabet +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions * `free_monoid α`: free monoid over alphabet `α`; defined as a synonym for `list α` diff --git a/src/algebra/module/hom.lean b/src/algebra/module/hom.lean index 1bc931a85ecce..81b6c29f0738e 100644 --- a/src/algebra/module/hom.lean +++ b/src/algebra/module/hom.lean @@ -8,6 +8,9 @@ import algebra.module.pi /-! # Bundled hom instances for module and multiplicative actions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines instances for module, mul_action and related structures on bundled `_hom` types. These are analogous to the instances in `algebra.module.pi`, but for bundled instead of unbundled diff --git a/src/algebra/module/pi.lean b/src/algebra/module/pi.lean index 6f5da6bf84e49..c91004eb8ebe1 100644 --- a/src/algebra/module/pi.lean +++ b/src/algebra/module/pi.lean @@ -11,6 +11,9 @@ import group_theory.group_action.pi /-! # Pi instances for modules +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines instances for module and related structures on Pi Types -/ diff --git a/src/algebra/order/archimedean.lean b/src/algebra/order/archimedean.lean index 0212343cf1717..3fa825a3a4457 100644 --- a/src/algebra/order/archimedean.lean +++ b/src/algebra/order/archimedean.lean @@ -9,6 +9,9 @@ import data.rat.floor /-! # Archimedean groups and fields. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the archimedean property for ordered groups and proves several results connected to this notion. Being archimedean means that for all elements `x` and `y>0` there exists a natural number `n` such that `x ≤ n • y`. diff --git a/src/algebra/order/floor.lean b/src/algebra/order/floor.lean index cf18e656a950e..498e145f7f46b 100644 --- a/src/algebra/order/floor.lean +++ b/src/algebra/order/floor.lean @@ -13,6 +13,9 @@ import tactic.positivity /-! # Floor and ceil +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Summary We define the natural- and integer-valued floor and ceil functions on linearly ordered rings. diff --git a/src/algebra/punit_instances.lean b/src/algebra/punit_instances.lean index a5b75fee4bb52..7b86493d96fd1 100644 --- a/src/algebra/punit_instances.lean +++ b/src/algebra/punit_instances.lean @@ -13,6 +13,9 @@ import order.complete_boolean_algebra /-! # Instances on punit +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file collects facts about algebraic structures on the one-element type, e.g. that it is a commutative ring. -/ diff --git a/src/algebra/star/basic.lean b/src/algebra/star/basic.lean index b9e18e02a0c72..6b74729fcc6b4 100644 --- a/src/algebra/star/basic.lean +++ b/src/algebra/star/basic.lean @@ -12,6 +12,9 @@ import data.set_like.basic /-! # Star monoids, rings, and modules +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We introduce the basic algebraic notions of star monoids, star rings, and star modules. A star algebra is simply a star ring that is also a star module. diff --git a/src/algebra/star/pi.lean b/src/algebra/star/pi.lean index b7d59f4906bfe..3aafeb402df34 100644 --- a/src/algebra/star/pi.lean +++ b/src/algebra/star/pi.lean @@ -9,6 +9,9 @@ import algebra.ring.pi /-! # `star` on pi types +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We put a `has_star` structure on pi types that operates elementwise, such that it describes the complex conjugation of vectors. -/ diff --git a/src/algebra/star/prod.lean b/src/algebra/star/prod.lean index a719f7948a483..3ced92aefa10a 100644 --- a/src/algebra/star/prod.lean +++ b/src/algebra/star/prod.lean @@ -10,6 +10,9 @@ import algebra.module.prod /-! # `star` on product types +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We put a `has_star` structure on product types that operates elementwise. -/ diff --git a/src/algebra/star/unitary.lean b/src/algebra/star/unitary.lean index 5d320f3b1c16b..e06dc567dd283 100644 --- a/src/algebra/star/unitary.lean +++ b/src/algebra/star/unitary.lean @@ -9,6 +9,9 @@ import group_theory.submonoid.operations /-! # Unitary elements of a star monoid +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines `unitary R`, where `R` is a star monoid, as the submonoid made of the elements that satisfy `star U * U = 1` and `U * star U = 1`, and these form a group. This includes, for instance, unitary operators on Hilbert spaces. diff --git a/src/data/list/big_operators/basic.lean b/src/data/list/big_operators/basic.lean index ac866d2ce39f3..738d87ea3f9e6 100644 --- a/src/data/list/big_operators/basic.lean +++ b/src/data/list/big_operators/basic.lean @@ -8,6 +8,9 @@ import data.list.forall2 /-! # Sums and products from lists +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file provides basic results about `list.prod`, `list.sum`, which calculate the product and sum of elements of a list and `list.alternating_prod`, `list.alternating_sum`, their alternating counterparts. These are defined in [`data.list.defs`](./defs). diff --git a/src/data/list/big_operators/lemmas.lean b/src/data/list/big_operators/lemmas.lean index e42d4982a4244..520d2cbba6d99 100644 --- a/src/data/list/big_operators/lemmas.lean +++ b/src/data/list/big_operators/lemmas.lean @@ -15,7 +15,10 @@ import algebra.ring.commute import data.int.units import data.set.basic -/-! # Lemmas about `list.sum` and `list.prod` requiring extra algebra imports -/ +/-! # Lemmas about `list.sum` and `list.prod` requiring extra algebra imports + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4.-/ open mul_opposite list diff --git a/src/data/list/chain.lean b/src/data/list/chain.lean index 1bd5fa52aa60c..5396aecfe3b4e 100644 --- a/src/data/list/chain.lean +++ b/src/data/list/chain.lean @@ -9,6 +9,9 @@ import logic.relation /-! # Relation chain +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file provides basic results about `list.chain` (definition in `data.list.defs`). A list `[a₂, ..., aₙ]` is a `chain` starting at `a₁` with respect to the relation `r` if `r a₁ a₂` and `r a₂ a₃` and ... and `r aₙ₋₁ aₙ`. We write it `chain r a₁ [a₂, ..., aₙ]`. diff --git a/src/data/list/count.lean b/src/data/list/count.lean index 1f77bc6d461d4..7332515c38d4f 100644 --- a/src/data/list/count.lean +++ b/src/data/list/count.lean @@ -8,6 +8,9 @@ import data.list.big_operators.basic /-! # Counting in lists +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves basic properties of `list.countp` and `list.count`, which count the number of elements of a list satisfying a predicate and equal to a given element respectively. Their definitions can be found in [`data.list.defs`](./defs). diff --git a/src/data/list/dedup.lean b/src/data/list/dedup.lean index a0a7ee3667055..b65f1122ff147 100644 --- a/src/data/list/dedup.lean +++ b/src/data/list/dedup.lean @@ -8,6 +8,9 @@ import data.list.nodup /-! # Erasure of duplicates in a list +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves basic results about `list.dedup` (definition in `data.list.defs`). `dedup l` returns `l` without its duplicates. It keeps the earliest (that is, rightmost) occurrence of each. diff --git a/src/data/list/duplicate.lean b/src/data/list/duplicate.lean index 638c67ce66c5f..efa5f35a83d93 100644 --- a/src/data/list/duplicate.lean +++ b/src/data/list/duplicate.lean @@ -8,6 +8,9 @@ import data.list.nodup /-! # List duplicates +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions * `list.duplicate x l : Prop` is an inductive property that holds when `x` is a duplicate in `l` diff --git a/src/data/list/join.lean b/src/data/list/join.lean index 590b5102b8cda..eaeca717c8225 100644 --- a/src/data/list/join.lean +++ b/src/data/list/join.lean @@ -8,6 +8,9 @@ import data.list.big_operators.basic /-! # Join of a list of lists +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves basic properties of `list.join`, which concatenates a list of lists. It is defined in [`data.list.defs`](./defs). -/ diff --git a/src/data/list/lattice.lean b/src/data/list/lattice.lean index e7c2b28cc3ea0..63d95e20fa561 100644 --- a/src/data/list/lattice.lean +++ b/src/data/list/lattice.lean @@ -11,6 +11,9 @@ import algebra.order.monoid.min_max /-! # Lattice structure of lists +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This files prove basic properties about `list.disjoint`, `list.union`, `list.inter` and `list.bag_inter`, which are defined in core Lean and `data.list.defs`. diff --git a/src/data/list/nodup.lean b/src/data/list/nodup.lean index ce702ff59712f..fc3b413305d01 100644 --- a/src/data/list/nodup.lean +++ b/src/data/list/nodup.lean @@ -11,6 +11,9 @@ import data.set.pairwise /-! # Lists with no duplicates +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + `list.nodup` is defined in `data/list/defs`. In this file we prove various properties of this predicate. -/ diff --git a/src/data/list/pairwise.lean b/src/data/list/pairwise.lean index 484f2fa4dee07..e702acbb2c163 100644 --- a/src/data/list/pairwise.lean +++ b/src/data/list/pairwise.lean @@ -11,6 +11,9 @@ import logic.relation /-! # Pairwise relations on a list +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file provides basic results about `list.pairwise` and `list.pw_filter` (definitions are in `data.list.defs`). `pairwise r [a 0, ..., a (n - 1)]` means `∀ i j, i < j → r (a i) (a j)`. For example, diff --git a/src/data/list/permutation.lean b/src/data/list/permutation.lean index 982f019d7cae1..d42d45bd24b58 100644 --- a/src/data/list/permutation.lean +++ b/src/data/list/permutation.lean @@ -8,6 +8,9 @@ import data.list.join /-! # Permutations of a list +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove properties about `list.permutations`, a list of all permutations of a list. It is defined in `data.list.defs`. diff --git a/src/data/list/prod_sigma.lean b/src/data/list/prod_sigma.lean index c720d8312cb34..da1eba10681bc 100644 --- a/src/data/list/prod_sigma.lean +++ b/src/data/list/prod_sigma.lean @@ -8,6 +8,9 @@ import data.list.big_operators.basic /-! # Lists in product and sigma types +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves basic properties of `list.product` and `list.sigma`, which are list constructions living in `prod` and `sigma` types respectively. Their definitions can be found in [`data.list.defs`](./defs). Beware, this is not about `list.prod`, the multiplicative product. diff --git a/src/data/list/range.lean b/src/data/list/range.lean index 04e0bcf17b944..94e3d326fd7bd 100644 --- a/src/data/list/range.lean +++ b/src/data/list/range.lean @@ -9,6 +9,9 @@ import data.list.zip /-! # Ranges of naturals as lists +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file shows basic results about `list.iota`, `list.range`, `list.range'` (all defined in `data.list.defs`) and defines `list.fin_range`. `fin_range n` is the list of elements of `fin n`. diff --git a/src/data/list/rdrop.lean b/src/data/list/rdrop.lean index be776a0524591..836372335cadc 100644 --- a/src/data/list/rdrop.lean +++ b/src/data/list/rdrop.lean @@ -11,6 +11,9 @@ import data.list.infix # Dropping or taking from lists on the right +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Taking or removing element from the tail end of a list ## Main defintions diff --git a/src/data/list/zip.lean b/src/data/list/zip.lean index 7202ba95e1501..1528725e6a2b3 100644 --- a/src/data/list/zip.lean +++ b/src/data/list/zip.lean @@ -9,6 +9,9 @@ import algebra.order.monoid.min_max /-! # zip & unzip +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file provides results about `list.zip_with`, `list.zip` and `list.unzip` (definitions are in core Lean). `zip_with f l₁ l₂` applies `f : α → β → γ` pointwise to a list `l₁ : list α` and `l₂ : list β`. It diff --git a/src/data/nat/bitwise.lean b/src/data/nat/bitwise.lean index 730cffbe8a5b6..54358c6ec377d 100644 --- a/src/data/nat/bitwise.lean +++ b/src/data/nat/bitwise.lean @@ -10,6 +10,9 @@ import tactic.linarith /-! # Bitwise operations on natural numbers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In the first half of this file, we provide theorems for reasoning about natural numbers from their bitwise properties. In the second half of this file, we show properties of the bitwise operations `lor`, `land` and `lxor`, which are defined in core. diff --git a/src/data/rat/floor.lean b/src/data/rat/floor.lean index 5428c18f5740e..28d4ce7a7e78b 100644 --- a/src/data/rat/floor.lean +++ b/src/data/rat/floor.lean @@ -10,6 +10,9 @@ import tactic.field_simp /-! # Floor Function for Rational Numbers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Summary We define the `floor` function and the `floor_ring` instance on `ℚ`. Some technical lemmas relating diff --git a/src/data/real/cau_seq.lean b/src/data/real/cau_seq.lean index 1fd077feeb150..de87799d093b0 100644 --- a/src/data/real/cau_seq.lean +++ b/src/data/real/cau_seq.lean @@ -13,6 +13,9 @@ import group_theory.group_action.pi /-! # Cauchy sequences +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A basic theory of Cauchy sequences, used in the construction of the reals and p-adic numbers. Where applicable, lemmas that will be reused in other contexts have been stated in extra generality. diff --git a/src/data/set/pointwise/basic.lean b/src/data/set/pointwise/basic.lean index df2bae5a87c48..9cd15e343d3fe 100644 --- a/src/data/set/pointwise/basic.lean +++ b/src/data/set/pointwise/basic.lean @@ -12,6 +12,9 @@ import data.nat.order.basic /-! # Pointwise operations of sets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines pointwise algebraic operations on sets. ## Main declarations diff --git a/src/logic/equiv/local_equiv.lean b/src/logic/equiv/local_equiv.lean index 901d263f3cdd2..7b1c693559b8f 100644 --- a/src/logic/equiv/local_equiv.lean +++ b/src/logic/equiv/local_equiv.lean @@ -9,6 +9,9 @@ import logic.equiv.defs /-! # Local equivalences +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This files defines equivalences between subsets of given types. An element `e` of `local_equiv α β` is made of two maps `e.to_fun` and `e.inv_fun` respectively from α to β and from β to α (just like equivs), which are inverse to each other on the subsets From fb987e222ccfe26760bd9410c036da5b488f92c1 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Wed, 11 Jan 2023 12:26:06 +0000 Subject: [PATCH 0208/1291] chore(analysis/seminorm): golf (#18089) A part of this PR is forward-ported as leanprover-community/mathlib4#1429 --- src/algebra/order/field/basic.lean | 4 ++ src/analysis/locally_convex/abs_convex.lean | 25 ++++------- src/analysis/locally_convex/bounded.lean | 2 +- .../locally_convex/with_seminorms.lean | 2 +- src/analysis/seminorm.lean | 44 +++++-------------- 5 files changed, 25 insertions(+), 52 deletions(-) diff --git a/src/algebra/order/field/basic.lean b/src/algebra/order/field/basic.lean index f7a94a50a75d3..126e30e1cd340 100644 --- a/src/algebra/order/field/basic.lean +++ b/src/algebra/order/field/basic.lean @@ -445,6 +445,10 @@ begin exact lt_max_iff.2 (or.inl $ lt_add_one _) end +lemma exists_pos_lt_mul {a : α} (h : 0 < a) (b : α) : ∃ c : α, 0 < c ∧ b < c * a := +let ⟨c, hc₀, hc⟩ := exists_pos_mul_lt h b +in ⟨c⁻¹, inv_pos.2 hc₀, by rwa [← div_eq_inv_mul, lt_div_iff hc₀]⟩ + lemma monotone.div_const {β : Type*} [preorder β] {f : β → α} (hf : monotone f) {c : α} (hc : 0 ≤ c) : monotone (λ x, (f x) / c) := begin diff --git a/src/analysis/locally_convex/abs_convex.lean b/src/analysis/locally_convex/abs_convex.lean index cfefd68a51df4..5377b4e2ad376 100644 --- a/src/analysis/locally_convex/abs_convex.lean +++ b/src/analysis/locally_convex/abs_convex.lean @@ -144,29 +144,22 @@ variables [smul_comm_class ℝ 𝕜 E] [locally_convex_space ℝ E] lemma with_gauge_seminorm_family : with_seminorms (gauge_seminorm_family 𝕜 E) := begin refine seminorm_family.with_seminorms_of_has_basis _ _, - refine filter.has_basis.to_has_basis (nhds_basis_abs_convex_open 𝕜 E) (λ s hs, _) (λ s hs, _), + refine (nhds_basis_abs_convex_open 𝕜 E).to_has_basis (λ s hs, _) (λ s hs, _), { refine ⟨s, ⟨_, rfl.subset⟩⟩, - rw seminorm_family.basis_sets_iff, - refine ⟨{⟨s, hs⟩}, 1, one_pos, _⟩, - simp only [finset.sup_singleton], - rw gauge_seminorm_family_ball, - simp only [subtype.coe_mk] }, + convert (gauge_seminorm_family _ _).basis_sets_singleton_mem ⟨s, hs⟩ one_pos, + rw [gauge_seminorm_family_ball, subtype.coe_mk] }, refine ⟨s, ⟨_, rfl.subset⟩⟩, rw seminorm_family.basis_sets_iff at hs, - rcases hs with ⟨t, r, hr, hs⟩, - rw seminorm.ball_finset_sup_eq_Inter _ _ _ hr at hs, - rw hs, + rcases hs with ⟨t, r, hr, rfl⟩, + rw [seminorm.ball_finset_sup_eq_Inter _ _ _ hr], -- We have to show that the intersection contains zero, is open, balanced, and convex refine ⟨mem_Inter₂.mpr (λ _ _, by simp [seminorm.mem_ball_zero, hr]), - is_open_bInter (to_finite _) (λ _ _, _), + is_open_bInter (to_finite _) (λ S _, _), balanced_Inter₂ (λ _ _, seminorm.balanced_ball_zero _ _), convex_Inter₂ (λ _ _, seminorm.convex_ball _ _ _)⟩, -- The only nontrivial part is to show that the ball is open have hr' : r = ‖(r : 𝕜)‖ * 1 := by simp [abs_of_pos hr], - have hr'' : (r : 𝕜) ≠ 0 := by simp [ne_of_gt hr], - rw hr', - rw ←seminorm.smul_ball_zero (norm_pos_iff.mpr hr''), - refine is_open.smul₀ _ hr'', - rw gauge_seminorm_family_ball, - exact abs_convex_open_sets.coe_is_open _, + have hr'' : (r : 𝕜) ≠ 0 := by simp [hr.ne'], + rw [hr', ← seminorm.smul_ball_zero hr'', gauge_seminorm_family_ball], + exact S.coe_is_open.smul₀ hr'' end diff --git a/src/analysis/locally_convex/bounded.lean b/src/analysis/locally_convex/bounded.lean index a2b3d5c7767ba..5c26f900fd27c 100644 --- a/src/analysis/locally_convex/bounded.lean +++ b/src/analysis/locally_convex/bounded.lean @@ -280,7 +280,7 @@ begin rcases h (metric.ball_mem_nhds 0 zero_lt_one) with ⟨ρ, hρ, hρball⟩, rcases normed_field.exists_lt_norm 𝕜 ρ with ⟨a, ha⟩, specialize hρball a ha.le, - rw [← ball_norm_seminorm 𝕜 E, seminorm.smul_ball_zero (hρ.trans ha), + rw [← ball_norm_seminorm 𝕜 E, seminorm.smul_ball_zero (norm_pos_iff.1 $ hρ.trans ha), ball_norm_seminorm, mul_one] at hρball, exact ⟨‖a‖, hρball.trans metric.ball_subset_closed_ball⟩ }, { exact λ ⟨C, hC⟩, (is_vonN_bounded_closed_ball 𝕜 E C).subset hC } diff --git a/src/analysis/locally_convex/with_seminorms.lean b/src/analysis/locally_convex/with_seminorms.lean index 0294d199c1771..c53496a47375a 100644 --- a/src/analysis/locally_convex/with_seminorms.lean +++ b/src/analysis/locally_convex/with_seminorms.lean @@ -488,7 +488,7 @@ begin rcases h with ⟨r, hr, h⟩, cases normed_field.exists_lt_norm 𝕜 r with a ha, specialize h a (le_of_lt ha), - rw [seminorm.smul_ball_zero (lt_trans hr ha), mul_one] at h, + rw [seminorm.smul_ball_zero (norm_pos_iff.1 $ hr.trans ha), mul_one] at h, refine ⟨‖a‖, lt_trans hr ha, _⟩, intros x hx, specialize h hx, diff --git a/src/analysis/seminorm.lean b/src/analysis/seminorm.lean index e373ef838d830..e939177b3737f 100644 --- a/src/analysis/seminorm.lean +++ b/src/analysis/seminorm.lean @@ -780,51 +780,28 @@ section normed_field variables [normed_field 𝕜] [add_comm_group E] [module 𝕜 E] (p : seminorm 𝕜 E) {A B : set E} {a : 𝕜} {r : ℝ} {x : E} -lemma smul_ball_zero {p : seminorm 𝕜 E} {k : 𝕜} {r : ℝ} (hk : 0 < ‖k‖) : +lemma smul_ball_zero {p : seminorm 𝕜 E} {k : 𝕜} {r : ℝ} (hk : k ≠ 0) : k • p.ball 0 r = p.ball 0 (‖k‖ * r) := begin ext, - rw [set.mem_smul_set, seminorm.mem_ball_zero], - split; intro h, - { rcases h with ⟨y, hy, h⟩, - rw [←h, map_smul_eq_mul], - rw seminorm.mem_ball_zero at hy, - exact (mul_lt_mul_left hk).mpr hy }, - refine ⟨k⁻¹ • x, _, _⟩, - { rw [seminorm.mem_ball_zero, map_smul_eq_mul, norm_inv, ←(mul_lt_mul_left hk), - ←mul_assoc, ←(div_eq_mul_inv ‖k‖ ‖k‖), div_self (ne_of_gt hk), one_mul], - exact h}, - rw [←smul_assoc, smul_eq_mul, ←div_eq_mul_inv, div_self (norm_pos_iff.mp hk), one_smul], + rw [mem_smul_set_iff_inv_smul_mem₀ hk, p.mem_ball_zero, p.mem_ball_zero, map_smul_eq_mul, + norm_inv, ← div_eq_inv_mul, div_lt_iff (norm_pos_iff.2 hk), mul_comm] end lemma ball_zero_absorbs_ball_zero (p : seminorm 𝕜 E) {r₁ r₂ : ℝ} (hr₁ : 0 < r₁) : absorbs 𝕜 (p.ball 0 r₁) (p.ball 0 r₂) := begin - by_cases hr₂ : r₂ ≤ 0, - { rw ball_eq_emptyset p hr₂, exact absorbs_empty }, - rw [not_le] at hr₂, - rcases exists_between hr₁ with ⟨r, hr, hr'⟩, - refine ⟨r₂/r, div_pos hr₂ hr, _⟩, - simp_rw set.subset_def, - intros a ha x hx, - have ha' : 0 < ‖a‖ := lt_of_lt_of_le (div_pos hr₂ hr) ha, - rw [smul_ball_zero ha', p.mem_ball_zero], + rcases exists_pos_lt_mul hr₁ r₂ with ⟨r, hr₀, hr⟩, + refine ⟨r, hr₀, λ a ha x hx, _⟩, + rw [smul_ball_zero (norm_pos_iff.1 $ hr₀.trans_le ha), p.mem_ball_zero], rw p.mem_ball_zero at hx, - rw div_le_iff hr at ha, - exact hx.trans (lt_of_le_of_lt ha ((mul_lt_mul_left ha').mpr hr')), + exact hx.trans (hr.trans_le $ mul_le_mul_of_nonneg_right ha hr₁.le) end /-- Seminorm-balls at the origin are absorbent. -/ protected lemma absorbent_ball_zero (hr : 0 < r) : absorbent 𝕜 (ball p (0 : E) r) := -begin - rw absorbent_iff_nonneg_lt, - rintro x, - have hxr : 0 ≤ p x / r := by positivity, - refine ⟨p x/r, hxr, λ a ha, _⟩, - have ha₀ : 0 < ‖a‖ := hxr.trans_lt ha, - refine ⟨a⁻¹ • x, _, smul_inv_smul₀ (norm_pos_iff.1 ha₀) x⟩, - rwa [mem_ball_zero, map_smul_eq_mul, norm_inv, inv_mul_lt_iff ha₀, ←div_lt_iff hr], -end +absorbent_iff_forall_absorbs_singleton.2 $ λ x, (p.ball_zero_absorbs_ball_zero hr).mono_right $ + singleton_subset_iff.2 $ p.mem_ball_zero.2 $ lt_add_one _ /-- Closed seminorm-balls at the origin are absorbent. -/ protected lemma absorbent_closed_ball_zero (hr : 0 < r) : absorbent 𝕜 (closed_ball p (0 : E) r) := @@ -946,8 +923,7 @@ begin suffices : (p.restrict_scalars ℝ).ball 0 ε ∈ (𝓝 0 : filter E), { rwa seminorm.ball_zero_eq_preimage_ball at this }, have := (set_smul_mem_nhds_zero_iff hε.ne.symm).mpr hp, - rwa [seminorm.smul_ball_zero (norm_pos_iff.mpr hε.ne.symm), - real.norm_of_nonneg hε.le, mul_one] at this + rwa [seminorm.smul_ball_zero hε.ne', real.norm_of_nonneg hε.le, mul_one] at this end protected lemma uniform_continuous_of_continuous_at_zero [uniform_space E] [uniform_add_group E] From 0bd2ea37bcba5769e14866170f251c9bc64e35d7 Mon Sep 17 00:00:00 2001 From: Chris Hughes Date: Wed, 11 Jan 2023 15:23:04 +0000 Subject: [PATCH 0209/1291] refactor(data/fin/basic): fin refactor to use ne_zero (#18107) Co-authored-by: Chris Hughes <33847686+ChrisHughes24@users.noreply.github.com> Co-authored-by: Riccardo Brasca --- src/algebraic_topology/Moore_complex.lean | 4 +- src/data/bitvec/basic.lean | 2 +- src/data/complex/module.lean | 7 +- src/data/fin/basic.lean | 186 +++++++++++++--------- src/data/sym/card.lean | 2 +- src/data/zmod/defs.lean | 25 ++- src/linear_algebra/matrix/adjugate.lean | 9 +- src/order/category/NonemptyFinLinOrd.lean | 2 +- src/order/jordan_holder.lean | 2 +- src/ring_theory/adjoin_root.lean | 5 +- src/ring_theory/chain_of_divisors.lean | 4 +- src/tactic/norm_fin.lean | 59 +++---- src/topology/locally_constant/basic.lean | 13 +- test/fin_cases.lean | 7 +- test/instance_diamonds.lean | 2 +- test/norm_fin.lean | 2 +- 16 files changed, 184 insertions(+), 147 deletions(-) diff --git a/src/algebraic_topology/Moore_complex.lean b/src/algebraic_topology/Moore_complex.lean index faf5988d0c504..79808b2f721a1 100644 --- a/src/algebraic_topology/Moore_complex.lean +++ b/src/algebraic_topology/Moore_complex.lean @@ -91,12 +91,12 @@ begin -- after the first simp the proofs are almost identical cases n; dsimp, { simp only [subobject.factor_thru_arrow_assoc], - slice_lhs 2 3 { erw ←X.δ_comp_δ (fin.zero_le 0), }, + slice_lhs 2 3 { erw ←X.δ_comp_δ (fin.zero_le (0 : fin (0 + 2))), }, rw ←factor_thru_arrow _ _ (finset_inf_arrow_factors finset.univ _ (0 : fin 2) (by simp)), slice_lhs 2 3 { rw [kernel_subobject_arrow_comp], }, simp, }, { simp [factor_thru_right], - slice_lhs 2 3 { erw ←X.δ_comp_δ (fin.zero_le 0), }, + slice_lhs 2 3 { erw ←X.δ_comp_δ (fin.zero_le (0 : fin (n.succ + 2))) }, rw ←factor_thru_arrow _ _ (finset_inf_arrow_factors finset.univ _ (0 : fin (n+3)) (by simp)), slice_lhs 2 3 { rw [kernel_subobject_arrow_comp], }, simp, }, diff --git a/src/data/bitvec/basic.lean b/src/data/bitvec/basic.lean index abc0d961fc3cd..c5715812f85be 100644 --- a/src/data/bitvec/basic.lean +++ b/src/data/bitvec/basic.lean @@ -77,7 +77,7 @@ begin end lemma to_fin_val {n : ℕ} (v : bitvec n) : (to_fin v : ℕ) = v.to_nat := -by rw [to_fin, fin.coe_of_nat_eq_mod', nat.mod_eq_of_lt]; apply to_nat_lt +by rw [to_fin, fin.of_nat'_eq_coe, fin.coe_of_nat_eq_mod', nat.mod_eq_of_lt]; apply to_nat_lt lemma to_fin_le_to_fin_of_le {n} {v₀ v₁ : bitvec n} (h : v₀ ≤ v₁) : v₀.to_fin ≤ v₁.to_fin := show (v₀.to_fin : ℕ) ≤ v₁.to_fin, diff --git a/src/data/complex/module.lean b/src/data/complex/module.lean index e5287b5b3742b..a79d585495917 100644 --- a/src/data/complex/module.lean +++ b/src/data/complex/module.lean @@ -148,10 +148,9 @@ basis.of_equiv_fun @[simp] lemma coe_basis_one_I : ⇑basis_one_I = ![1, I] := funext $ λ i, basis.apply_eq_iff.mpr $ finsupp.ext $ λ j, by fin_cases i; fin_cases j; - simp only [coe_basis_one_I_repr, finsupp.single_eq_same, finsupp.single_eq_of_ne, - matrix.cons_val_zero, matrix.cons_val_one, matrix.head_cons, - nat.one_ne_zero, fin.one_eq_zero_iff, fin.zero_eq_one_iff, ne.def, not_false_iff, - one_re, one_im, I_re, I_im] + simp only [coe_basis_one_I_repr, finsupp.single_eq_of_ne, matrix.cons_val_zero, + matrix.cons_val_one, matrix.head_cons, fin.one_eq_zero_iff, ne.def, not_false_iff, I_re, + nat.succ_succ_ne_one, one_im, I_im, one_re, finsupp.single_eq_same, fin.zero_eq_one_iff] instance : finite_dimensional ℝ ℂ := of_fintype_basis basis_one_I diff --git a/src/data/fin/basic.lean b/src/data/fin/basic.lean index 2fe63657d1da1..aae2f9ffddc75 100644 --- a/src/data/fin/basic.lean +++ b/src/data/fin/basic.lean @@ -251,18 +251,28 @@ def factorial {n : ℕ} : fin n → ℕ instance {n : ℕ} : has_well_founded (fin n) := ⟨_, measure_wf coe⟩ -@[simp] lemma coe_zero {n : ℕ} : ((0 : fin (n+1)) : ℕ) = 0 := rfl +instance has_zero_of_ne_zero [ne_zero n] : has_zero (fin n) := +⟨⟨0, ne_zero.pos _⟩⟩ + +/-- Given a positive `n`, `fin.of_nat' i` is `i % n` as an element of `fin n`. -/ +def of_nat' [ne_zero n] (i : ℕ) : fin n := ⟨i%n, mod_lt _ $ ne_zero.pos n⟩ + +instance has_one_of_ne_zero [ne_zero n] : has_one (fin n) := +⟨of_nat' 1⟩ + +@[simp] lemma coe_zero (n : ℕ) [ne_zero n] : ((0 : fin n) : ℕ) = 0 := rfl attribute [simp] val_zero -@[simp] lemma val_zero' (n) : (0 : fin (n+1)).val = 0 := rfl -@[simp] lemma mk_zero : (⟨0, nat.succ_pos'⟩ : fin (n + 1)) = (0 : fin _) := rfl +@[simp] lemma val_zero' (n) [ne_zero n] : (0 : fin n).val = 0 := rfl +@[simp] lemma mk_zero [ne_zero n] : + (⟨0, nat.pos_of_ne_zero (ne_zero.ne n)⟩ : fin n) = (0 : fin _) := rfl -@[simp] lemma zero_le (a : fin (n + 1)) : 0 ≤ a := zero_le a.1 +@[simp] lemma zero_le [ne_zero n] (a : fin n) : 0 ≤ a := zero_le a.1 lemma zero_lt_one : (0 : fin (n + 2)) < 1 := nat.zero_lt_one @[simp] lemma not_lt_zero (a : fin n.succ) : ¬a < 0. -lemma pos_iff_ne_zero (a : fin (n+1)) : 0 < a ↔ a ≠ 0 := +lemma pos_iff_ne_zero [ne_zero n] (a : fin n) : 0 < a ↔ a ≠ 0 := by rw [← coe_fin_lt, coe_zero, pos_iff_ne_zero, ne.def, ne.def, ext_iff, coe_zero] lemma eq_zero_or_eq_succ {n : ℕ} (i : fin (n+1)) : i = 0 ∨ ∃ j : fin n, i = j.succ := @@ -323,7 +333,7 @@ lemma last_val (n : ℕ) : (last n).val = n := rfl theorem le_last (i : fin (n+1)) : i ≤ last n := le_of_lt_succ i.is_lt -instance : bounded_order (fin (n + 1)) := +instance : bounded_order (fin (n+1)) := { top := last n, le_top := le_last, bot := 0, @@ -339,7 +349,7 @@ le_antisymm (le_last i) (not_lt.1 h) lemma top_eq_last (n : ℕ) : ⊤ = fin.last n := rfl -lemma bot_eq_zero (n : ℕ) : ⊥ = (0 : fin (n + 1)) := rfl +lemma bot_eq_zero (n : ℕ) : ⊥ = (0 : fin (n+1)) := rfl section @@ -394,13 +404,10 @@ section add ### addition, numerals, and coercion from nat -/ -/-- Given a positive `n`, `fin.of_nat' i` is `i % n` as an element of `fin n`. -/ -def of_nat' [ne_zero n] (i : ℕ) : fin n := ⟨i%n, mod_lt _ $ ne_zero.pos n⟩ - lemma one_val {n : ℕ} : (1 : fin (n+1)).val = 1 % (n+1) := rfl -lemma coe_one' {n : ℕ} : ((1 : fin (n+1)) : ℕ) = 1 % (n+1) := rfl -@[simp] lemma val_one {n : ℕ} : (1 : fin (n+2)).val = 1 := rfl -@[simp] lemma coe_one {n : ℕ} : ((1 : fin (n+2)) : ℕ) = 1 := rfl +lemma coe_one' (n : ℕ) [ne_zero n] : ((1 : fin n) : ℕ) = 1 % n := rfl +@[simp] lemma val_one (n : ℕ) : (1 : fin (n+2)).val = 1 := rfl +@[simp] lemma coe_one (n : ℕ) : ((1 : fin (n+2)) : ℕ) = 1 := rfl @[simp] lemma mk_one : (⟨1, nat.succ_lt_succ (nat.succ_pos n)⟩ : fin (n + 2)) = (1 : fin _) := rfl instance {n : ℕ} : nontrivial (fin (n + 2)) := ⟨⟨0, 1, dec_trivial⟩⟩ @@ -413,13 +420,13 @@ by rcases n with _|_|n; simp [is_empty.subsingleton, unique.subsingleton, not_su section monoid -@[simp] protected lemma add_zero (k : fin (n + 1)) : k + 0 = k := +@[simp] protected lemma add_zero [ne_zero n] (k : fin n) : k + 0 = k := by simp [eq_iff_veq, add_def, mod_eq_of_lt (is_lt k)] -@[simp] protected lemma zero_add (k : fin (n + 1)) : (0 : fin (n + 1)) + k = k := +@[simp] protected lemma zero_add [ne_zero n] (k : fin n) : (0 : fin n) + k = k := by simp [eq_iff_veq, add_def, mod_eq_of_lt (is_lt k)] -instance add_comm_monoid (n : ℕ) : add_comm_monoid (fin (n + 1)) := +instance add_comm_monoid (n : ℕ) [ne_zero n] : add_comm_monoid (fin n) := { add := (+), add_assoc := by simp [eq_iff_veq, add_def, add_assoc], zero := 0, @@ -427,9 +434,9 @@ instance add_comm_monoid (n : ℕ) : add_comm_monoid (fin (n + 1)) := add_zero := fin.add_zero, add_comm := by simp [eq_iff_veq, add_def, add_comm] } -instance : add_monoid_with_one (fin (n + 1)) := +instance [ne_zero n] : add_monoid_with_one (fin n) := { one := 1, - nat_cast := fin.of_nat, + nat_cast := fin.of_nat', nat_cast_zero := rfl, nat_cast_succ := λ i, eq_of_veq (add_mod _ _ _), .. fin.add_comm_monoid n } @@ -450,12 +457,9 @@ by rw [fin.coe_add, nat.add_mod_eq_ite, lemma coe_bit0 {n : ℕ} (k : fin n) : ((bit0 k : fin n) : ℕ) = bit0 (k : ℕ) % n := by { cases k, refl } -lemma coe_bit1 {n : ℕ} (k : fin (n + 1)) : - ((bit1 k : fin (n + 1)) : ℕ) = bit1 (k : ℕ) % (n + 1) := -begin - cases n, { cases k with k h, cases k, {show _ % _ = _, simp}, cases h with _ h, cases h }, - simp [bit1, fin.coe_bit0, fin.coe_add, fin.coe_one], -end +lemma coe_bit1 {n : ℕ} [ne_zero n] (k : fin n) : + ((bit1 k : fin n) : ℕ) = bit1 (k : ℕ) % n := +by simp [bit1, coe_add, coe_bit0, coe_one'] lemma coe_add_one_of_lt {n : ℕ} {i : fin n.succ} (h : i < last _) : (↑(i + 1) : ℕ) = i + 1 := @@ -486,8 +490,8 @@ section bit (⟨bit0 m, h⟩ : fin n) = (bit0 ⟨m, (nat.le_add_right m m).trans_lt h⟩ : fin _) := eq_of_veq (nat.mod_eq_of_lt h).symm -@[simp] lemma mk_bit1 {m n : ℕ} (h : bit1 m < n + 1) : - (⟨bit1 m, h⟩ : fin (n + 1)) = (bit1 ⟨m, (nat.le_add_right m m).trans_lt +@[simp] lemma mk_bit1 {m n : ℕ} [ne_zero n] (h : bit1 m < n) : + (⟨bit1 m, h⟩ : fin n) = (bit1 ⟨m, (nat.le_add_right m m).trans_lt ((m + m).lt_succ_self.trans h)⟩ : fin _) := begin ext, @@ -506,32 +510,36 @@ section of_nat_coe lemma of_nat_eq_coe (n : ℕ) (a : ℕ) : (of_nat a : fin (n+1)) = a := rfl -/-- Converting an in-range number to `fin (n + 1)` produces a result +@[simp] +lemma of_nat'_eq_coe (n : ℕ) [ne_zero n] (a : ℕ) : (of_nat' a : fin n) = a := +rfl + +/-- Converting an in-range number to `fin n` produces a result whose value is the original number. -/ -lemma coe_val_of_lt {n : ℕ} {a : ℕ} (h : a < n + 1) : - ((a : fin (n + 1)).val) = a := +lemma coe_val_of_lt {n : ℕ} [ne_zero n] {a : ℕ} (h : a < n) : + ((a : fin n).val) = a := begin - rw ←of_nat_eq_coe, + rw ←of_nat'_eq_coe, exact nat.mod_eq_of_lt h end -/-- Converting the value of a `fin (n + 1)` to `fin (n + 1)` results +/-- Converting the value of a `fin n` to `fin n` results in the same value. -/ -lemma coe_val_eq_self {n : ℕ} (a : fin (n + 1)) : (a.val : fin (n + 1)) = a := +lemma coe_val_eq_self {n : ℕ} [ne_zero n] (a : fin n) : (a.val : fin n) = a := begin rw fin.eq_iff_veq, exact coe_val_of_lt a.property end -/-- Coercing an in-range number to `fin (n + 1)`, and converting back +/-- Coercing an in-range number to `fin n`, and converting back to `ℕ`, results in that number. -/ -lemma coe_coe_of_lt {n : ℕ} {a : ℕ} (h : a < n + 1) : - ((a : fin (n + 1)) : ℕ) = a := +lemma coe_coe_of_lt {n : ℕ} [ne_zero n] {a : ℕ} (h : a < n) : + ((a : fin n) : ℕ) = a := coe_val_of_lt h -/-- Converting a `fin (n + 1)` to `ℕ` and back results in the same +/-- Converting a `fin n` to `ℕ` and back results in the same value. -/ -@[simp] lemma coe_coe_eq_self {n : ℕ} (a : fin (n + 1)) : ((a : ℕ) : fin (n + 1)) = a := +@[simp] lemma coe_coe_eq_self {n : ℕ} [ne_zero n] (a : fin n) : ((a : ℕ) : fin n) = a := coe_val_eq_self a lemma coe_nat_eq_last (n) : (n : fin (n + 1)) = fin.last n := @@ -555,16 +563,17 @@ lemma one_pos : (0 : fin (n + 2)) < 1 := succ_pos 0 lemma zero_ne_one : (0 : fin (n + 2)) ≠ 1 := ne_of_lt one_pos -@[simp] lemma zero_eq_one_iff : (0 : fin (n + 1)) = 1 ↔ n = 0 := +@[simp] lemma zero_eq_one_iff [ne_zero n]: (0 : fin n) = 1 ↔ n = 1 := begin split, - { cases n; intro h, - { refl }, - { have := zero_ne_one, contradiction } }, - { rintro rfl, refl } + { intro h, + have := congr_arg (coe : fin n → ℕ) h, + simp only [fin.coe_zero, ← nat.dvd_iff_mod_eq_zero, fin.coe_one', @eq_comm _ 0] at this, + exact eq_one_of_dvd_one this }, + { unfreezingI { rintro rfl }, refl } end -@[simp] lemma one_eq_zero_iff : (1 : fin (n + 1)) = 0 ↔ n = 0 := +@[simp] lemma one_eq_zero_iff [ne_zero n]: (1 : fin n) = 0 ↔ n = 1 := by rw [eq_comm, zero_eq_one_iff] end add @@ -602,9 +611,25 @@ lemma succ_injective (n : ℕ) : injective (@fin.succ n) := lemma succ_ne_zero {n} : ∀ k : fin n, fin.succ k ≠ 0 | ⟨k, hk⟩ heq := nat.succ_ne_zero k $ ext_iff.1 heq -@[simp] lemma succ_zero_eq_one : fin.succ (0 : fin (n + 1)) = 1 := rfl +@[simp] lemma succ_zero_eq_one [ne_zero n] : fin.succ (0 : fin n) = 1 := +begin + unfreezingI { cases n }, + { exact (ne_zero.ne 0 rfl).elim }, + { refl } +end + +/-- Version of `succ_zero_eq_one` to be used by `dsimp` -/ +@[simp] lemma succ_zero_eq_one' : fin.succ (0 : fin (n+1)) = 1 := rfl + +@[simp] lemma succ_one_eq_two [ne_zero n] : fin.succ (1 : fin (n + 1)) = 2 := +begin + unfreezingI { cases n }, + { exact (ne_zero.ne 0 rfl).elim }, + { refl } +end -@[simp] lemma succ_one_eq_two : fin.succ (1 : fin (n + 2)) = 2 := rfl +/-- Version of `succ_one_eq_two` to be used by `dsimp` -/ +@[simp] lemma succ_one_eq_two' : fin.succ (1 : fin (n + 2)) = 2 := rfl @[simp] lemma succ_mk (n i : ℕ) (h : i < n) : fin.succ ⟨i, h⟩ = ⟨i + 1, nat.succ_lt_succ h⟩ := rfl @@ -652,9 +677,9 @@ begin simp end -@[simp] lemma le_zero_iff {n : ℕ} {k : fin (n + 1)} : +@[simp] lemma le_zero_iff {n : ℕ} [ne_zero n] {k : fin n} : k ≤ 0 ↔ k = 0 := -le_bot_iff +⟨λ h, fin.eq_of_veq $ by rw [nat.eq_zero_of_le_zero h]; refl, by rintro rfl; refl⟩ lemma succ_succ_ne_one (a : fin n) : fin.succ (fin.succ a) ≠ 1 := ne_of_gt (one_lt_succ_succ a) @@ -713,8 +738,8 @@ as it is eligible for `dsimp`. -/ @[simp] lemma coe_cast (h : n = m) (i : fin n) : (cast h i : ℕ) = i := rfl -@[simp] lemma cast_zero {n' : ℕ} {h : n + 1 = n' + 1} : - cast h (0 : fin (n + 1)) = 0 := +@[simp] lemma cast_zero {n' : ℕ} [ne_zero n] {h : n = n'} : + cast h (0 : fin n) = by { haveI : ne_zero n' := by { rw ← h; apply_instance }, exact 0 } := ext rfl @[simp] lemma cast_last {n' : ℕ} {h : n + 1 = n' + 1} : @@ -836,18 +861,18 @@ lemma cast_succ_inj {a b : fin n} : a.cast_succ = b.cast_succ ↔ a = b := lemma cast_succ_lt_last (a : fin n) : cast_succ a < last n := lt_iff_coe_lt_coe.mpr a.is_lt -@[simp] lemma cast_succ_zero : cast_succ (0 : fin (n + 1)) = 0 := rfl +@[simp] lemma cast_succ_zero [ne_zero n] : cast_succ (0 : fin n) = 0 := rfl @[simp] lemma cast_succ_one {n : ℕ} : fin.cast_succ (1 : fin (n + 2)) = 1 := rfl /-- `cast_succ i` is positive when `i` is positive -/ -lemma cast_succ_pos {i : fin (n + 1)} (h : 0 < i) : 0 < cast_succ i := +lemma cast_succ_pos [ne_zero n] {i : fin n} (h : 0 < i) : 0 < cast_succ i := by simpa [lt_iff_coe_lt_coe] using h -@[simp] lemma cast_succ_eq_zero_iff (a : fin (n + 1)) : a.cast_succ = 0 ↔ a = 0 := +@[simp] lemma cast_succ_eq_zero_iff [ne_zero n] (a : fin n) : a.cast_succ = 0 ↔ a = 0 := fin.ext_iff.trans $ (fin.ext_iff.trans $ by exact iff.rfl).symm -lemma cast_succ_ne_zero_iff (a : fin (n + 1)) : a.cast_succ ≠ 0 ↔ a ≠ 0 := +lemma cast_succ_ne_zero_iff [ne_zero n] (a : fin n) : a.cast_succ ≠ 0 ↔ a ≠ 0 := not_iff_not.mpr $ cast_succ_eq_zero_iff a lemma cast_succ_fin_succ (n : ℕ) (j : fin n) : @@ -874,7 +899,7 @@ by { rw [cast_succ, lt_iff_coe_lt_coe, coe_cast_add, coe_succ], exact lt_add_one set.range (cast_succ : fin n → fin n.succ) = {i | (i : ℕ) < n} := range_cast_le _ -@[simp] lemma coe_of_injective_cast_succ_symm {n : ℕ} (i : fin n.succ) (hi) : +@[simp] lemma coe_of_injective_cast_succ_symm {n : ℕ} (i : fin (n+1)) (hi) : ((equiv.of_injective cast_succ (cast_succ_injective _)).symm ⟨i, hi⟩ : ℕ) = i := begin rw ← coe_cast_succ, @@ -1325,15 +1350,15 @@ open nat int /-- Negation on `fin n` -/ instance (n : ℕ) : has_neg (fin n) := ⟨λ a, ⟨(n - a) % n, nat.mod_lt _ a.pos⟩⟩ -/-- Abelian group structure on `fin (n+1)`. -/ -instance (n : ℕ) : add_comm_group (fin (n+1)) := +/-- Abelian group structure on `fin n`. -/ +instance (n : ℕ) [ne_zero n] : add_comm_group (fin n) := { add_left_neg := λ ⟨a, ha⟩, fin.ext $ trans (nat.mod_add_mod _ _ _) $ by { rw [fin.coe_mk, fin.coe_zero, tsub_add_cancel_of_le, nat.mod_self], exact le_of_lt ha }, sub_eq_add_neg := λ ⟨a, ha⟩ ⟨b, hb⟩, fin.ext $ - show (a + (n + 1 - b)) % (n + 1) = (a + (n + 1 - b) % (n + 1)) % (n + 1), by simp, + show (a + (n - b)) % n = (a + (n - b) % n) % n, by simp, sub := fin.sub, ..fin.add_comm_monoid n, - ..fin.has_neg n.succ } + ..fin.has_neg n } protected lemma coe_neg (a : fin n) : ((-a : fin n) : ℕ) = (n - a) % n := rfl @@ -1434,18 +1459,19 @@ lemma succ_above_below (p : fin (n + 1)) (i : fin n) (h : i.cast_succ < p) : p.succ_above i = i.cast_succ := by { rw [succ_above], exact if_pos h } -@[simp] lemma succ_above_ne_zero_zero {a : fin (n + 2)} (ha : a ≠ 0) : a.succ_above 0 = 0 := +@[simp] lemma succ_above_ne_zero_zero [ne_zero n] {a : fin (n + 1)} (ha : a ≠ 0) : + a.succ_above 0 = 0 := begin rw fin.succ_above_below, { refl }, { exact bot_lt_iff_ne_bot.mpr ha } end -lemma succ_above_eq_zero_iff {a : fin (n + 2)} {b : fin (n + 1)} (ha : a ≠ 0) : +lemma succ_above_eq_zero_iff [ne_zero n] {a : fin (n + 1)} {b : fin n} (ha : a ≠ 0) : a.succ_above b = 0 ↔ b = 0 := by simp only [←succ_above_ne_zero_zero ha, order_embedding.eq_iff_eq] -lemma succ_above_ne_zero {a : fin (n + 2)} {b : fin (n + 1)} (ha : a ≠ 0) (hb : b ≠ 0) : +lemma succ_above_ne_zero [ne_zero n] {a : fin (n + 1)} {b : fin n} (ha : a ≠ 0) (hb : b ≠ 0) : a.succ_above b ≠ 0 := mt (succ_above_eq_zero_iff ha).mp hb @@ -1516,7 +1542,8 @@ begin end /-- Embedding a positive `fin n` results in a positive fin (n + 1)` -/ -lemma succ_above_pos (p : fin (n + 2)) (i : fin (n + 1)) (h : 0 < i) : 0 < p.succ_above i := +lemma succ_above_pos [ne_zero n] (p : fin (n + 1)) (i : fin n) (h : 0 < i) : + 0 < p.succ_above i := begin by_cases H : i.cast_succ < p, { simpa [succ_above_below _ _ H] using cast_succ_pos h }, @@ -1588,7 +1615,7 @@ succ_above_left_injective.eq_iff (0 : fin (n + 1)).succ_above i = i.succ := rfl -@[simp] lemma succ_succ_above_zero {n : ℕ} (i : fin (n + 1)) : +@[simp] lemma succ_succ_above_zero {n : ℕ} [ne_zero n] (i : fin n) : (i.succ).succ_above 0 = 0 := succ_above_below _ _ (succ_pos _) @@ -1606,9 +1633,9 @@ succ_succ_above_zero 0 /-- By moving `succ` to the outside of this expression, we create opportunities for further simplification using `succ_above_zero` or `succ_succ_above_zero`. -/ -@[simp] lemma succ_succ_above_one {n : ℕ} (i : fin (n + 2)) : +@[simp] lemma succ_succ_above_one {n : ℕ} [ne_zero n] (i : fin (n + 1)) : (i.succ).succ_above 1 = (i.succ_above 0).succ := -succ_succ_above_succ i 0 +by rw [← succ_succ_above_succ i 0, succ_zero_eq_one] @[simp] lemma one_succ_above_succ {n : ℕ} (j : fin n) : (1 : fin (n + 2)).succ_above j.succ = j.succ.succ := @@ -1839,15 +1866,15 @@ def clamp (n m : ℕ) : fin (m + 1) := of_nat $ min n m @[simp] lemma coe_clamp (n m : ℕ) : (clamp n m : ℕ) = min n m := nat.mod_eq_of_lt $ nat.lt_succ_iff.mpr $ min_le_right _ _ +@[simp] lemma coe_of_nat_eq_mod' (m n : ℕ) [I : ne_zero m] : + ((n : fin m) : ℕ) = n % m := +rfl + @[simp] lemma coe_of_nat_eq_mod (m n : ℕ) : ((n : fin (succ m)) : ℕ) = n % succ m := by rw [← of_nat_eq_coe]; refl -@[simp] lemma coe_of_nat_eq_mod' (m n : ℕ) [I : ne_zero m] : - (@fin.of_nat' _ I n : ℕ) = n % m := -rfl - section mul /-! @@ -1860,16 +1887,25 @@ lemma val_mul {n : ℕ} : ∀ a b : fin n, (a * b).val = (a.val * b.val) % n lemma coe_mul {n : ℕ} : ∀ a b : fin n, ((a * b : fin n) : ℕ) = (a * b) % n | ⟨_, _⟩ ⟨_, _⟩ := rfl -@[simp] protected lemma mul_one (k : fin (n + 1)) : k * 1 = k := -by { cases n, simp, simp [eq_iff_veq, mul_def, mod_eq_of_lt (is_lt k)] } +@[simp] protected lemma mul_one [ne_zero n] (k : fin n) : k * 1 = k := +begin + unfreezingI { cases n }, + { simp }, + unfreezingI { cases n }, + { simp }, + simp [eq_iff_veq, mul_def, mod_eq_of_lt (is_lt k)] +end + +protected lemma mul_comm (a b : fin n) : a * b = b * a := +fin.eq_of_veq $ by rw [mul_def, mul_def, mul_comm] -@[simp] protected lemma one_mul (k : fin (n + 1)) : (1 : fin (n + 1)) * k = k := -by { cases n, simp, simp [eq_iff_veq, mul_def, mod_eq_of_lt (is_lt k)] } +@[simp] protected lemma one_mul [ne_zero n] (k : fin n) : (1 : fin n) * k = k := +by rw [fin.mul_comm, fin.mul_one] -@[simp] protected lemma mul_zero (k : fin (n + 1)) : k * 0 = 0 := +@[simp] protected lemma mul_zero [ne_zero n] (k : fin n) : k * 0 = 0 := by simp [eq_iff_veq, mul_def] -@[simp] protected lemma zero_mul (k : fin (n + 1)) : (0 : fin (n + 1)) * k = 0 := +@[simp] protected lemma zero_mul [ne_zero n] (k : fin n) : (0 : fin n) * k = 0 := by simp [eq_iff_veq, mul_def] end mul diff --git a/src/data/sym/card.lean b/src/data/sym/card.lean index ce49ebe29e323..213c8d23cb071 100644 --- a/src/data/sym/card.lean +++ b/src/data/sym/card.lean @@ -90,7 +90,7 @@ protected def E2 {n k : ℕ} : { simp only [fin.zero_succ_above, map_map, comp_app], nth_rewrite_rhs 0 ←(map_id' s), refine sym.map_congr (λ v hv, _), - rw [←fin.zero_succ_above v, ←fin.cast_succ_zero, fin.pred_above_succ_above 0 v] } } + rw [←fin.zero_succ_above v, ←@fin.cast_succ_zero n.succ, fin.pred_above_succ_above 0 v] } } lemma card_sym_fin_eq_multichoose (n k : ℕ) : card (sym (fin n) k) = multichoose n k := begin diff --git a/src/data/zmod/defs.lean b/src/data/zmod/defs.lean index b89419abb1e48..3a91f1a2e5f37 100644 --- a/src/data/zmod/defs.lean +++ b/src/data/zmod/defs.lean @@ -38,25 +38,24 @@ to register the ring structure on `zmod n` as type class instance. open nat.modeq int -/-- Multiplicative commutative semigroup structure on `fin (n+1)`. -/ -instance (n : ℕ) : comm_semigroup (fin (n+1)) := +/-- Multiplicative commutative semigroup structure on `fin n`. -/ +instance (n : ℕ) : comm_semigroup (fin n) := { mul_assoc := λ ⟨a, ha⟩ ⟨b, hb⟩ ⟨c, hc⟩, fin.eq_of_veq - (calc ((a * b) % (n+1) * c) ≡ a * b * c [MOD (n+1)] : (nat.mod_modeq _ _).mul_right _ - ... ≡ a * (b * c) [MOD (n+1)] : by rw mul_assoc - ... ≡ a * (b * c % (n+1)) [MOD (n+1)] : (nat.mod_modeq _ _).symm.mul_left _), - mul_comm := λ ⟨a, _⟩ ⟨b, _⟩, - fin.eq_of_veq (show (a * b) % (n+1) = (b * a) % (n+1), by rw mul_comm), + (calc ((a * b) % n * c) ≡ a * b * c [MOD n] : (nat.mod_modeq _ _).mul_right _ + ... ≡ a * (b * c) [MOD n] : by rw mul_assoc + ... ≡ a * (b * c % n) [MOD n] : (nat.mod_modeq _ _).symm.mul_left _), + mul_comm := fin.mul_comm, ..fin.has_mul } -private lemma left_distrib_aux (n : ℕ) : ∀ a b c : fin (n+1), a * (b + c) = a * b + a * c := +private lemma left_distrib_aux (n : ℕ) : ∀ a b c : fin n, a * (b + c) = a * b + a * c := λ ⟨a, ha⟩ ⟨b, hb⟩ ⟨c, hc⟩, fin.eq_of_veq -(calc a * ((b + c) % (n+1)) ≡ a * (b + c) [MOD (n+1)] : (nat.mod_modeq _ _).mul_left _ - ... ≡ a * b + a * c [MOD (n+1)] : by rw mul_add - ... ≡ (a * b) % (n+1) + (a * c) % (n+1) [MOD (n+1)] : +(calc a * ((b + c) % n) ≡ a * (b + c) [MOD n] : (nat.mod_modeq _ _).mul_left _ + ... ≡ a * b + a * c [MOD n] : by rw mul_add + ... ≡ (a * b) % n + (a * c) % n [MOD n] : (nat.mod_modeq _ _).symm.add (nat.mod_modeq _ _).symm) -/-- Commutative ring structure on `fin (n+1)`. -/ -instance (n : ℕ) : comm_ring (fin (n+1)) := +/-- Commutative ring structure on `fin n`. -/ +instance (n : ℕ) [ne_zero n] : comm_ring (fin n) := { one_mul := fin.one_mul, mul_one := fin.mul_one, left_distrib := left_distrib_aux n, diff --git a/src/linear_algebra/matrix/adjugate.lean b/src/linear_algebra/matrix/adjugate.lean index 749fb6023e554..9652c44789d95 100644 --- a/src/linear_algebra/matrix/adjugate.lean +++ b/src/linear_algebra/matrix/adjugate.lean @@ -344,9 +344,12 @@ begin ext i j, rw [adjugate_apply, det_fin_two], fin_cases i; fin_cases j; - simp only [nat.one_ne_zero, one_mul, fin.one_eq_zero_iff, pi.single_eq_same, zero_mul, - fin.zero_eq_one_iff, sub_zero, pi.single_eq_of_ne, ne.def, not_false_iff, update_row_self, - update_row_ne, cons_val_zero, mul_zero, mul_one, zero_sub, cons_val_one, head_cons, of_apply], + simp only [one_mul, fin.one_eq_zero_iff, pi.single_eq_same, mul_zero, sub_zero, + pi.single_eq_of_ne, ne.def, not_false_iff, update_row_self, update_row_ne, cons_val_zero, + of_apply, nat.succ_succ_ne_one, pi.single_eq_of_ne, update_row_self, pi.single_eq_of_ne, ne.def, + fin.zero_eq_one_iff, nat.succ_succ_ne_one, not_false_iff, update_row_ne, fin.one_eq_zero_iff, + zero_mul, pi.single_eq_same, one_mul, zero_sub, of_apply, cons_val', cons_val_fin_one, + cons_val_one, head_fin_const, neg_inj, eq_self_iff_true, cons_val_zero, head_cons, mul_one] end @[simp] lemma adjugate_fin_two_of (a b c d : α) : diff --git a/src/order/category/NonemptyFinLinOrd.lean b/src/order/category/NonemptyFinLinOrd.lean index fb2ff89921896..95a42944534e9 100644 --- a/src/order/category/NonemptyFinLinOrd.lean +++ b/src/order/category/NonemptyFinLinOrd.lean @@ -132,7 +132,7 @@ begin { exfalso, exact h₂ (le_of_lt h₁), }, { exfalso, exact hm a (eq_of_le_of_not_lt h₂ h₁), }, }, simpa only [order_hom.coe_fun_mk, lt_self_iff_false, if_false, le_refl, if_true, - ulift.up_inj, fin.one_eq_zero_iff] using h, }, + ulift.up_inj, fin.one_eq_zero_iff, nat.succ_succ_ne_one] using h, }, { intro h, exact concrete_category.epi_of_surjective f h, }, end diff --git a/src/order/jordan_holder.lean b/src/order/jordan_holder.lean index ab622f412faae..fade2bc333fae 100644 --- a/src/order/jordan_holder.lean +++ b/src/order/jordan_holder.lean @@ -522,7 +522,7 @@ fin.snoc_cast_succ _ _ _ @[simp] lemma bot_snoc (s : composition_series X) (x : X) (hsat : is_maximal s.top x) : (snoc s x hsat).bot = s.bot := -by rw [bot, bot, ← fin.cast_succ_zero, snoc_cast_succ] +by rw [bot, bot, ← snoc_cast_succ s _ _ 0, fin.cast_succ_zero] lemma mem_snoc {s : composition_series X} {x y: X} {hsat : is_maximal s.top x} : y ∈ snoc s x hsat ↔ y ∈ s ∨ y = x := diff --git a/src/ring_theory/adjoin_root.lean b/src/ring_theory/adjoin_root.lean index 08f9229b1d320..a99085c2611d8 100644 --- a/src/ring_theory/adjoin_root.lean +++ b/src/ring_theory/adjoin_root.lean @@ -301,7 +301,7 @@ lemma mk_surjective (hg : g.monic) : function.surjective (mk g) := /-- The elements `1, root g, ..., root g ^ (d - 1)` form a basis for `adjoin_root g`, where `g` is a monic polynomial of degree `d`. -/ -@[simps] def power_basis_aux' (hg : g.monic) : +def power_basis_aux' (hg : g.monic) : basis (fin g.nat_degree) R (adjoin_root g) := basis.of_equiv_fun { to_fun := λ f i, (mod_by_monic_hom hg f).coeff i, @@ -328,6 +328,9 @@ basis.of_equiv_fun exact j.2 } }, end} +-- This was moved after the definition to prevent a timeout +attribute [simps] power_basis_aux' + /-- The power basis `1, root g, ..., root g ^ (d - 1)` for `adjoin_root g`, where `g` is a monic polynomial of degree `d`. -/ @[simps] def power_basis' (hg : g.monic) : power_basis R (adjoin_root g) := diff --git a/src/ring_theory/chain_of_divisors.lean b/src/ring_theory/chain_of_divisors.lean index 115c19ce86355..a64688573412a 100644 --- a/src/ring_theory/chain_of_divisors.lean +++ b/src/ring_theory/chain_of_divisors.lean @@ -110,7 +110,7 @@ begin { contradiction }, obtain ⟨i, rfl⟩ := h₂.1 (dvd_trans hp' hr), refine congr_arg c (eq_of_ge_of_not_gt _ $ λ hi, _), - { rw [fin.le_iff_coe_le_coe, fin.coe_one, nat.succ_le_iff, ← fin.coe_zero, + { rw [fin.le_iff_coe_le_coe, fin.coe_one, nat.succ_le_iff, ← fin.coe_zero (n.succ + 1), ← fin.lt_iff_coe_lt_coe, fin.pos_iff_ne_zero], rintro rfl, exact hp.not_unit (first_of_chain_is_unit h₁ @h₂) }, @@ -119,7 +119,7 @@ begin refine not_irreducible_of_not_unit_dvd_not_unit (dvd_not_unit.not_unit (associates.dvd_not_unit_iff_lt.2 (h₁ (show (0 : fin (n + 2)) < j, from _)) )) _ hp.irreducible, - { simpa [← fin.succ_zero_eq_one, fin.succ_lt_succ_iff] using hi }, + { simpa [fin.succ_lt_succ_iff, fin.lt_iff_coe_lt_coe] using hi }, { refine associates.dvd_not_unit_iff_lt.2 (h₁ _), simpa only [fin.coe_eq_cast_succ] using fin.lt_succ } end diff --git a/src/tactic/norm_fin.lean b/src/tactic/norm_fin.lean index 80b4d74941df7..2136b060c2eab 100644 --- a/src/tactic/norm_fin.lean +++ b/src/tactic/norm_fin.lean @@ -56,9 +56,10 @@ by rw ← h.coe; exact a.2 theorem normalize_fin_lt.of {n a b} (h : normalize_fin_lt n a b) : normalize_fin n a b := h.trans $ eq.symm $ nat.mod_eq_of_lt h.lt -theorem normalize_fin.zero (n) : normalize_fin (n+1) 0 0 := by { rw normalize_fin, norm_num } -theorem normalize_fin_lt.zero (n) : normalize_fin_lt (n+1) 0 0 := refl _ -theorem normalize_fin.one (n) : normalize_fin (n+1) 1 1 := refl _ +theorem normalize_fin.zero (n : ℕ) [ne_zero n] : + normalize_fin n 0 0 := by { rw normalize_fin, norm_num } +theorem normalize_fin_lt.zero (n : ℕ) [ne_zero n] : normalize_fin_lt n 0 0 := refl _ +theorem normalize_fin.one (n : ℕ) [ne_zero n] : normalize_fin n 1 1 := refl _ theorem normalize_fin.add {n} {a b : fin n} {a' b' c' : ℕ} (ha : normalize_fin n a a') (hb : normalize_fin n b b') (h : a' + b' = c') : normalize_fin n (a + b) c' := @@ -69,8 +70,8 @@ theorem normalize_fin.mul {n} {a b : fin n} {a' b' c' : ℕ} by simp only [normalize_fin, ← h] at *; rw [nat.mul_mod, ← ha, ← hb, fin.mul_def] theorem normalize_fin.bit0 {n} {a : fin n} {a' : ℕ} (h : normalize_fin n a a') : normalize_fin n (bit0 a) (bit0 a') := h.add h rfl -theorem normalize_fin.bit1 {n} {a : fin (n+1)} {a' : ℕ} - (h : normalize_fin (n+1) a a') : normalize_fin (n+1) (bit1 a) (bit1 a') := +theorem normalize_fin.bit1 {n : ℕ} [ne_zero n] {a : fin n} {a' : ℕ} + (h : normalize_fin n a a') : normalize_fin n (bit1 a) (bit1 a') := h.bit0.add (normalize_fin.one _) rfl theorem normalize_fin_lt.succ {n} {a : fin n} {a' b : ℕ} @@ -180,20 +181,20 @@ do ic ← mk_instance_cache `(ℕ), (a, _) ← state_t.run m (ic, none), pure a direct expr pattern match because expr pattern matches generate very large terms under the hood so going via an intermediate inductive type like this is more efficient. -/ meta inductive match_fin_result -| zero (n : expr) -- `(0 : fin (n+1))` -| one (n : expr) -- `(1 : fin (n+1))` -| add (n a b : expr) -- `(a + b : fin n)` -| mul (n a b : expr) -- `(a * b : fin n)` -| bit0 (n a : expr) -- `(bit0 a : fin n)` -| bit1 (n a : expr) -- `(bit1 a : fin (n+1))` -| succ (n a : expr) -- `(fin.succ a : fin n.succ)` -| cast_lt (n m i h : expr) -- `(fin.cast_lt (i : fin m) (h : i.val < n) : fin n)` -| cast_le (n m h a : expr) -- `(fin.cast_le (h : n ≤ m) (a : fin n) : fin m)` -| cast (n m h a : expr) -- `(fin.cast_le (h : n = m) (a : fin n) : fin m)` -| cast_add (n m a : expr) -- `(fin.cast_add m (a : fin n) : fin (n + m))` -| cast_succ (n a : expr) -- `(fin.cast_succ (a : fin n) : fin (n + 1))` -| add_nat (n m a : expr) -- `(fin.add_nat m (a : fin n) : fin (n + m))` -| nat_add (n m a : expr) -- `(fin.nat_add n (a : fin m) : fin (n + m))` +| zero (n : expr) (n0 : expr) -- `(0 : fin n)` +| one (n : expr) (n0 : expr) -- `(1 : fin n)` +| add (n a b : expr) -- `(a + b : fin n)` +| mul (n a b : expr) -- `(a * b : fin n)` +| bit0 (n a : expr) -- `(bit0 a : fin n)` +| bit1 (n a : expr) (n0 : expr) -- `(bit1 a : fin n)` +| succ (n a : expr) -- `(fin.succ a : fin n.succ)` +| cast_lt (n m i h : expr) -- `(fin.cast_lt (i : fin m) (h : i.val < n) : fin n)` +| cast_le (n m h a : expr) -- `(fin.cast_le (h : n ≤ m) (a : fin n) : fin m)` +| cast (n m h a : expr) -- `(fin.cast_le (h : n = m) (a : fin n) : fin m)` +| cast_add (n m a : expr) -- `(fin.cast_add m (a : fin n) : fin (n + m))` +| cast_succ (n a : expr) -- `(fin.cast_succ (a : fin n) : fin (n + 1))` +| add_nat (n m a : expr) -- `(fin.add_nat m (a : fin n) : fin (n + m))` +| nat_add (n m a : expr) -- `(fin.nat_add n (a : fin m) : fin (n + m))` section open match_fin_result @@ -212,12 +213,12 @@ meta def match_fin_coe_fn (a : expr) : expr → option match_fin_result /-- Match a fin expression to a `match_fin_result`, for easier pattern matching in the evaluator. -/ meta def match_fin : expr → option match_fin_result -| `(@has_zero.zero ._ (@fin.has_zero %%n)) := some (zero n) -| `(@has_one.one ._ (@fin.has_one %%n)) := some (one n) +| `(@has_zero.zero ._ (@fin.has_zero_of_ne_zero %%n %%n0)) := some (zero n n0) +| `(@has_one.one ._ (@fin.has_one_of_ne_zero %%n %%n0)) := some (one n n0) | `(@has_add.add (fin %%n) ._ %%a %%b) := some (add n a b) | `(@has_mul.mul (fin %%n) ._ %%a %%b) := some (mul n a b) | `(@_root_.bit0 (fin %%n) ._ %%a) := some (bit0 n a) -| `(@_root_.bit1 ._ (@fin.has_one %%n) ._ %%a) := some (bit1 n a) +| `(@_root_.bit1 ._ (@fin.has_one_of_ne_zero %%n %%n0) ._ %%a) := some (bit1 n a n0) | `(@fin.succ %%n %%a) := some (succ n a) | `(@fin.cast_lt %%n %%m %%a %%h) := some (cast_lt n m a h) | (expr.app `(@coe_fn ._ ._ ._ %%f) a) := match_fin_coe_fn a f @@ -306,8 +307,8 @@ meta def eval_fin : expr → eval_fin_m (expr × expr) | a := do m ← match_fin a, match m with - | match_fin_result.zero n := pure (`(0 : ℕ), `(normalize_fin.zero).mk_app [n]) - | match_fin_result.one n := pure (`(1 : ℕ), `(normalize_fin.one).mk_app [n]) + | match_fin_result.zero n n0 := pure (`(0 : ℕ), `(normalize_fin.zero).mk_app [n, n0]) + | match_fin_result.one n n0 := pure (`(1 : ℕ), `(normalize_fin.one).mk_app [n, n0]) | match_fin_result.add n a b := do (a', pa) ← eval_fin a, (b', pb) ← eval_fin b, @@ -321,9 +322,9 @@ meta def eval_fin : expr → eval_fin_m (expr × expr) | match_fin_result.bit0 n a := do (a', pa) ← eval_fin a, pure (`(@bit0 ℕ _).mk_app [a'], `(@normalize_fin.bit0).mk_app [n, a, a', pa]) - | match_fin_result.bit1 n a := do + | match_fin_result.bit1 n a n0 := do (a', pa) ← eval_fin a, - pure (`(@bit1 ℕ _ _).mk_app [a'], `(@normalize_fin.bit1).mk_app [n, a, a', pa]) + pure (`(@bit1 ℕ _ _).mk_app [a'], `(@normalize_fin.bit1).mk_app [n, n0, a, a', pa]) | match_fin_result.cast m n nm a := do (a', pa) ← (eval_fin a).reset, pure (a', `(@normalize_fin.cast).mk_app [n, m, nm, a, a', pa]) @@ -397,10 +398,10 @@ meta def mk_fin_numeral (n m : expr) : expr → option (expr × expr) | a := match match_numeral a with | zero := some ( expr.app `(@has_zero.zero (fin %%n)) `(@fin.has_zero %%m), - expr.app `(normalize_fin.zero) m) + `(normalize_fin.zero).mk_app [n, `(@ne_zero.succ %%m)]) | one := some ( expr.app `(@has_one.one (fin %%n)) `(@fin.has_one %%m), - expr.app `(normalize_fin.one) m) + `(normalize_fin.one).mk_app [n, `(@ne_zero.succ %%m)]) | bit0 a := do (a', p) ← mk_fin_numeral a, some (`(bit0 %%a' : fin %%n), `(@normalize_fin.bit0).mk_app [n, a', a, p]) @@ -408,7 +409,7 @@ meta def mk_fin_numeral (n m : expr) : expr → option (expr × expr) (a', p) ← mk_fin_numeral a, some ( `(@_root_.bit1 (fin %%n)).mk_app [`(@fin.has_one %%m), `(@fin.has_add %%n), a'], - `(@normalize_fin.bit1).mk_app [m, a', a, p]) + `(@normalize_fin.bit1).mk_app [n, `(@ne_zero.succ %%m), a', a, p]) | _ := none end end diff --git a/src/topology/locally_constant/basic.lean b/src/topology/locally_constant/basic.lean index 1148549c6bcd2..ffae47b95d1ed 100644 --- a/src/topology/locally_constant/basic.lean +++ b/src/topology/locally_constant/basic.lean @@ -293,8 +293,8 @@ def of_clopen {X : Type*} [topological_space X] {U : set X} [∀ x, decidable (x fin_cases e, { convert hU.1 using 1, ext, - simp only [nat.one_ne_zero, mem_singleton_iff, fin.one_eq_zero_iff, - mem_preimage, ite_eq_left_iff], + simp only [mem_singleton_iff, fin.one_eq_zero_iff, mem_preimage, ite_eq_left_iff, + nat.succ_succ_ne_one], tauto }, { rw ← is_closed_compl_iff, convert hU.2, @@ -306,8 +306,8 @@ def of_clopen {X : Type*} [topological_space X] {U : set X} [∀ x, decidable (x [∀ x, decidable (x ∈ U)] (hU : is_clopen U) : of_clopen hU ⁻¹' ({0} : set (fin 2)) = U := begin ext, - simp only [of_clopen, nat.one_ne_zero, mem_singleton_iff, - fin.one_eq_zero_iff, coe_mk, mem_preimage, ite_eq_left_iff], + simp only [of_clopen, mem_singleton_iff, fin.one_eq_zero_iff, coe_mk, mem_preimage, + ite_eq_left_iff, nat.succ_succ_ne_one], tauto, end @@ -315,9 +315,8 @@ end [∀ x, decidable (x ∈ U)] (hU : is_clopen U) : of_clopen hU ⁻¹' ({1} : set (fin 2)) = Uᶜ := begin ext, - simp only [of_clopen, nat.one_ne_zero, mem_singleton_iff, coe_mk, - fin.zero_eq_one_iff, mem_preimage, ite_eq_right_iff, - mem_compl_iff], + simp only [of_clopen, mem_singleton_iff, coe_mk, fin.zero_eq_one_iff, mem_preimage, + ite_eq_right_iff, mem_compl_iff, nat.succ_succ_ne_one], tauto, end diff --git a/test/fin_cases.lean b/test/fin_cases.lean index 043d4585cc8e2..eaadb04f31273 100644 --- a/test/fin_cases.lean +++ b/test/fin_cases.lean @@ -105,21 +105,18 @@ end In some circumstances involving `let`, the temporary hypothesis that `fin_cases` creates does not get deleted. We test that this is correctly named and that the name can be changed. - -Note: after `fin_cases`, we have `this : (a : fin 3) = (0 : fin (2 + 1))` -for some reason. I don't know why, and it complicates the test. -/ example (f : ℕ → fin 3) : true := begin let a := f 3, fin_cases a, guard_hyp a := f 3, - guard_hyp this : a = (0 : fin (2 + 1)), + guard_hyp this : a = (0 : fin 3), trivial, trivial, let b := f 2, fin_cases b using what, - guard_hyp what : b = (0 : fin (2 + 1)), + guard_hyp what : b = (0 : fin 3), all_goals {trivial} end diff --git a/test/instance_diamonds.lean b/test/instance_diamonds.lean index 9e0db87bd6486..8cdfb34b3d0d4 100644 --- a/test/instance_diamonds.lean +++ b/test/instance_diamonds.lean @@ -243,7 +243,7 @@ example : @euclidean_domain.to_comm_ring _ (@field.to_euclidean_domain _ (zmod.f zmod.comm_ring p := rfl -example (n : ℕ) : zmod.comm_ring (n + 1) = fin.comm_ring n := rfl +example (n : ℕ) : zmod.comm_ring (n + 1) = fin.comm_ring (n + 1) := rfl example : zmod.comm_ring 0 = int.comm_ring := rfl end zmod diff --git a/test/norm_fin.lean b/test/norm_fin.lean index 9888495cf6f4a..2651c4ce3dd0e 100644 --- a/test/norm_fin.lean +++ b/test/norm_fin.lean @@ -28,7 +28,7 @@ example : equiv.swap (0 : fin 3) 1 (fin.succ 1) = 2 := begin success_if_fail {guard_target ((equiv.swap (0 : fin 3) 1) 2 = 2)}, norm_fin, - guard_target (equiv.swap (0 : fin 3) 1 2 = 2), + guard_target_mod_implicit (equiv.swap (0 : fin 3) 1 2 = 2), exact equiv.swap_apply_of_ne_of_ne dec_trivial dec_trivial end example : equiv.swap (0 : fin (1 + 2)) (1 : fin (nat.succ (1 + 1))) (fin.succ 1) = 2 := From 4568629f486f9762caf4f0c36ebcbad4b9797289 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 11 Jan 2023 15:23:07 +0000 Subject: [PATCH 0210/1291] chore(linear_algebra/tensor_product): generalize `tensor_tensor_tensor_comm_symm_tmul` (#18110) This lemma can be stated unapplied, making it more powerful --- src/linear_algebra/tensor_product.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/linear_algebra/tensor_product.lean b/src/linear_algebra/tensor_product.lean index e4bc0279f7e81..49321d45a4bcd 100644 --- a/src/linear_algebra/tensor_product.lean +++ b/src/linear_algebra/tensor_product.lean @@ -788,8 +788,8 @@ variables {M N P Q} tensor_tensor_tensor_comm R M N P Q ((m ⊗ₜ n) ⊗ₜ (p ⊗ₜ q)) = (m ⊗ₜ p) ⊗ₜ (n ⊗ₜ q) := rfl -@[simp] lemma tensor_tensor_tensor_comm_symm_tmul (m : M) (n : N) (p : P) (q : Q) : - (tensor_tensor_tensor_comm R M N P Q).symm ((m ⊗ₜ p) ⊗ₜ (n ⊗ₜ q)) = (m ⊗ₜ n) ⊗ₜ (p ⊗ₜ q) := +@[simp] lemma tensor_tensor_tensor_comm_symm : + (tensor_tensor_tensor_comm R M N P Q).symm = tensor_tensor_tensor_comm R M P N Q := rfl variables (M N P Q) From 5758ed34939d01a6884d1ce5f2a45f9e4e606216 Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Wed, 11 Jan 2023 15:23:09 +0000 Subject: [PATCH 0211/1291] feat(data/fin/vec_notation): add matrix.range_singleton and matrix.range_pair (#18112) --- src/data/fin/vec_notation.lean | 7 +++++++ 1 file changed, 7 insertions(+) diff --git a/src/data/fin/vec_notation.lean b/src/data/fin/vec_notation.lean index 41a6acc587c32..e19013da44475 100644 --- a/src/data/fin/vec_notation.lean +++ b/src/data/fin/vec_notation.lean @@ -124,6 +124,13 @@ set.ext $ λ y, by simp [fin.exists_fin_succ, eq_comm] @[simp] lemma range_empty (u : fin 0 → α) : set.range u = ∅ := set.range_eq_empty _ +@[simp] lemma range_cons_empty (x : α) (u : fin 0 → α) : set.range (matrix.vec_cons x u) = {x} := +by rw [range_cons, range_empty, set.union_empty] + +@[simp] lemma range_cons_cons_empty (x y : α) (u : fin 0 → α) : + set.range (vec_cons x $ vec_cons y u) = {x, y} := +by rw [range_cons, range_cons_empty, set.singleton_union] + @[simp] lemma vec_cons_const (a : α) : vec_cons a (λ k : fin n, a) = λ _, a := funext $ fin.forall_fin_succ.2 ⟨rfl, cons_val_succ _ _⟩ From 7c523cb78f4153682c2929e3006c863bfef463d0 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Wed, 11 Jan 2023 18:25:14 +0000 Subject: [PATCH 0212/1291] feat(data/list/count): partially sync with multiset (#18125) * rename `list.count_repeat` to `list.count_repeat_self`; * prove `list.count_repeat` with `if .. then .. else` in the RHS. --- archive/100-theorems-list/30_ballot_problem.lean | 4 ++-- src/data/list/count.lean | 10 ++++++++-- src/data/multiset/basic.lean | 7 +------ 3 files changed, 11 insertions(+), 10 deletions(-) diff --git a/archive/100-theorems-list/30_ballot_problem.lean b/archive/100-theorems-list/30_ballot_problem.lean index 6e9e589531cce..f0d659d67a7aa 100644 --- a/archive/100-theorems-list/30_ballot_problem.lean +++ b/archive/100-theorems-list/30_ballot_problem.lean @@ -82,7 +82,7 @@ begin { rwa ← list.count_eq_length.2 this }, { exact λ x hx, (this x hx).symm } }, { rintro rfl, - simp only [mem_set_of_eq, list.count_repeat, eq_self_iff_true, true_and], + simp only [mem_set_of_eq, list.count_repeat_self, eq_self_iff_true, true_and], refine ⟨list.count_eq_zero_of_not_mem _, λ x, _⟩; rw list.mem_repeat, { norm_num }, { rintro ⟨-, rfl⟩, @@ -105,7 +105,7 @@ begin { rwa ← list.count_eq_length.2 this }, { exact λ x hx, (this x hx).symm } }, { rintro rfl, - simp only [mem_set_of_eq, list.count_repeat, eq_self_iff_true, true_and], + simp only [mem_set_of_eq, list.count_repeat_self, eq_self_iff_true, true_and], refine ⟨list.count_eq_zero_of_not_mem _, λ x, _⟩; rw list.mem_repeat, { norm_num }, { rintro ⟨-, rfl⟩, diff --git a/src/data/list/count.lean b/src/data/list/count.lean index 7332515c38d4f..a26f76e73b4c8 100644 --- a/src/data/list/count.lean +++ b/src/data/list/count.lean @@ -164,17 +164,23 @@ lemma not_mem_of_count_eq_zero {a : α} {l : list α} (h : count a l = 0) : a @[simp] lemma count_eq_length {a : α} {l} : count a l = l.length ↔ ∀ b ∈ l, a = b := countp_eq_length _ -@[simp] lemma count_repeat (a : α) (n : ℕ) : count a (repeat a n) = n := +@[simp] lemma count_repeat_self (a : α) (n : ℕ) : count a (repeat a n) = n := by rw [count, countp_eq_length_filter, filter_eq_self.2, length_repeat]; exact λ b m, (eq_of_mem_repeat m).symm +lemma count_repeat (a b : α) (n : ℕ) : count a (repeat b n) = if a = b then n else 0 := +begin + split_ifs with h, + exacts [h ▸ count_repeat_self _ _, count_eq_zero_of_not_mem (mt eq_of_mem_repeat h)] +end + lemma le_count_iff_repeat_sublist {a : α} {l : list α} {n : ℕ} : n ≤ count a l ↔ repeat a n <+ l := ⟨λ h, ((repeat_sublist_repeat a).2 h).trans $ have filter (eq a) l = repeat a (count a l), from eq_repeat.2 ⟨by simp only [count, countp_eq_length_filter], λ b m, (of_mem_filter m).symm⟩, by rw ← this; apply filter_sublist, - λ h, by simpa only [count_repeat] using h.count_le a⟩ + λ h, by simpa only [count_repeat_self] using h.count_le a⟩ lemma repeat_count_eq_of_count_eq_length {a : α} {l : list α} (h : count a l = length l) : repeat a (count a l) = l := diff --git a/src/data/multiset/basic.lean b/src/data/multiset/basic.lean index 1881abd1b3170..ed38dc1edbfe5 100644 --- a/src/data/multiset/basic.lean +++ b/src/data/multiset/basic.lean @@ -1790,12 +1790,7 @@ by simp [repeat] theorem count_repeat (a b : α) (n : ℕ) : count a (repeat b n) = if (a = b) then n else 0 := -begin - split_ifs with h₁, - { rw [h₁, count_repeat_self] }, - { rw [count_eq_zero], - apply mt eq_of_mem_repeat h₁ }, -end +count_repeat a b n @[simp] theorem count_erase_self (a : α) (s : multiset α) : count a (erase s a) = pred (count a s) := From 914cb1775729bab4ef05ccddf6dd349652dc44fd Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Wed, 11 Jan 2023 21:36:31 +0000 Subject: [PATCH 0213/1291] feat(analysis/special_functions): Gaussian integral with complex parameter (#18090) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Compute `∫ x:ℝ, exp (-b * x^2)` when `b` is complex (with positive imaginary part). --- src/analysis/special_functions/gaussian.lean | 254 +++++++++++++------ src/topology/algebra/field.lean | 69 +++++ src/topology/connected.lean | 14 + 3 files changed, 264 insertions(+), 73 deletions(-) diff --git a/src/analysis/special_functions/gaussian.lean b/src/analysis/special_functions/gaussian.lean index 870e4675fdd23..ef4792ec70e2d 100644 --- a/src/analysis/special_functions/gaussian.lean +++ b/src/analysis/special_functions/gaussian.lean @@ -5,11 +5,17 @@ Authors: Sébastien Gouëzel -/ import analysis.special_functions.gamma import analysis.special_functions.polar_coord +import analysis.convex.complex /-! # Gaussian integral -We prove the formula `∫ x, exp (-b * x^2) = sqrt (π / b)`, in `integral_gaussian`. +We prove various versions of the formula for the Gaussian integral: +* `integral_gaussian`: for real `b` we have `∫ x:ℝ, exp (-b * x^2) = sqrt (π / b)`. +* `integral_gaussian_complex`: for complex `b` with `0 < re b` we have + `∫ x:ℝ, exp (-b * x^2) = (π / b) ^ (1 / 2)`. +* `integral_gaussian_Ioi` and `integral_gaussian_complex_Ioi`: variants for integrals over `Ioi 0`. +* `complex.Gamma_one_half_eq`: the formula `Γ (1 / 2) = √π`. -/ noncomputable theory @@ -17,6 +23,9 @@ noncomputable theory open real set measure_theory filter asymptotics open_locale real topological_space +open complex (hiding exp continuous_exp abs_of_nonneg) +notation `cexp` := complex.exp + lemma exp_neg_mul_sq_is_o_exp_neg {b : ℝ} (hb : 0 < b) : (λ x:ℝ, exp (-b * x^2)) =o[at_top] (λ x:ℝ, exp (-x)) := begin @@ -76,11 +85,7 @@ end lemma integrable_exp_neg_mul_sq {b : ℝ} (hb : 0 < b) : integrable (λ x:ℝ, exp (-b * x^2)) := -begin - have A : (-1 : ℝ) < 0, by norm_num, - convert integrable_rpow_mul_exp_neg_mul_sq hb A, - simp, -end +by simpa using integrable_rpow_mul_exp_neg_mul_sq hb (by norm_num : (-1 : ℝ) < 0) lemma integrable_on_Ioi_exp_neg_mul_sq_iff {b : ℝ} : integrable_on (λ x:ℝ, exp (-b * x^2)) (Ioi 0) ↔ 0 < b := @@ -98,109 +103,212 @@ end lemma integrable_exp_neg_mul_sq_iff {b : ℝ} : integrable (λ x:ℝ, exp (-b * x^2)) ↔ 0 < b := ⟨λ h, integrable_on_Ioi_exp_neg_mul_sq_iff.mp h.integrable_on, integrable_exp_neg_mul_sq⟩ -lemma integrable_mul_exp_neg_mul_sq {b : ℝ} (hb : 0 < b) : - integrable (λ x:ℝ, x * exp (-b * x^2)) := +lemma integrable_mul_exp_neg_mul_sq {b : ℝ} (hb : 0 < b) : integrable (λ x:ℝ, x * exp (-b * x^2)) := +by simpa using integrable_rpow_mul_exp_neg_mul_sq hb (by norm_num : (-1 : ℝ) < 1) + +lemma norm_cexp_neg_mul_sq (b : ℂ) (x : ℝ) : ‖complex.exp (-b * x^2)‖ = exp (-b.re * x^2) := +by rw [complex.norm_eq_abs, complex.abs_exp, ←of_real_pow, mul_comm (-b) _, of_real_mul_re, + neg_re, mul_comm] + +lemma integrable_cexp_neg_mul_sq {b : ℂ} (hb : 0 < b.re) : integrable (λ x:ℝ, cexp (-b * x^2)) := +begin + refine ⟨(complex.continuous_exp.comp + (continuous_const.mul (continuous_of_real.pow 2))).ae_strongly_measurable, _⟩, + rw ←has_finite_integral_norm_iff, + simp_rw norm_cexp_neg_mul_sq, + exact (integrable_exp_neg_mul_sq hb).2, +end + +lemma integrable_mul_cexp_neg_mul_sq {b : ℂ} (hb : 0 < b.re) : + integrable (λ x:ℝ, ↑x * cexp (-b * x^2)) := begin - have A : (-1 : ℝ) < 1, by norm_num, - convert integrable_rpow_mul_exp_neg_mul_sq hb A, - simp, + refine ⟨(continuous_of_real.mul (complex.continuous_exp.comp _)).ae_strongly_measurable, _⟩, + { exact continuous_const.mul (continuous_of_real.pow 2) }, + have := (integrable_mul_exp_neg_mul_sq hb).has_finite_integral, + rw ←has_finite_integral_norm_iff at this ⊢, + convert this, + ext1 x, + rw [norm_mul, norm_mul, norm_cexp_neg_mul_sq b, complex.norm_eq_abs, abs_of_real, + real.norm_eq_abs, norm_of_nonneg (exp_pos _).le], end -lemma integral_mul_exp_neg_mul_sq {b : ℝ} (hb : 0 < b) : - ∫ r in Ioi 0, r * exp (-b * r ^ 2) = (2 * b)⁻¹ := +lemma integral_mul_cexp_neg_mul_sq {b : ℂ} (hb : 0 < b.re) : + ∫ r:ℝ in Ioi 0, (r : ℂ) * cexp (-b * r ^ 2) = (2 * b)⁻¹ := begin - have I : integrable (λ x, x * exp (-b * x^2)) := integrable_mul_exp_neg_mul_sq hb, - refine tendsto_nhds_unique - (interval_integral_tendsto_integral_Ioi _ I.integrable_on filter.tendsto_id) _, - have A : ∀ x, has_deriv_at (λ x, - (2 * b)⁻¹ * exp (-b * x^2)) (x * exp (- b * x^2)) x, - { assume x, - convert (((has_deriv_at_pow 2 x)).const_mul (-b)).exp.const_mul (- (2 * b)⁻¹) using 1, - field_simp [hb.ne'], + have hb' : b ≠ 0 := by { contrapose! hb, rw [hb, zero_re], }, + refine tendsto_nhds_unique (interval_integral_tendsto_integral_Ioi _ + (integrable_mul_cexp_neg_mul_sq hb).integrable_on filter.tendsto_id) _, + have A : ∀ x:ℂ, has_deriv_at (λ x, - (2 * b)⁻¹ * cexp (-b * x^2)) (x * cexp (- b * x^2)) x, + { intro x, + convert (((has_deriv_at_pow 2 x)).const_mul (-b)).cexp.const_mul (- (2 * b)⁻¹) using 1, + field_simp [hb'], ring }, - have : ∀ (y : ℝ), ∫ x in 0..(id y), x * exp (- b * x^2) - = (- (2 * b)⁻¹ * exp (-b * y^2)) - (- (2 * b)⁻¹ * exp (-b * 0^2)) := - λ y, interval_integral.integral_eq_sub_of_has_deriv_at (λ x hx, A x) I.interval_integrable, - simp_rw [this], - have L : tendsto (λ (x : ℝ), (2 * b)⁻¹ - (2 * b)⁻¹ * exp (-b * x ^ 2)) at_top + have : ∀ (y : ℝ), ∫ x in 0..(id y), ↑x * cexp (-b * x^2) + = (- (2 * b)⁻¹ * cexp (-b * y^2)) - (- (2 * b)⁻¹ * cexp (-b * 0^2)) := + λ y, interval_integral.integral_eq_sub_of_has_deriv_at + (λ x hx, (A x).comp_of_real) (integrable_mul_cexp_neg_mul_sq hb).interval_integrable, + simp_rw this, + have L : tendsto (λ (x : ℝ), (2 * b)⁻¹ - (2 * b)⁻¹ * cexp (-b * x ^ 2)) at_top (𝓝 ((2 * b)⁻¹ - (2 * b)⁻¹ * 0)), - { refine tendsto_const_nhds.sub _, - apply tendsto.const_mul, - apply tendsto_exp_at_bot.comp, - exact tendsto.neg_const_mul_at_top (neg_lt_zero.2 hb) (tendsto_pow_at_top two_ne_zero) }, + { refine tendsto_const_nhds.sub (tendsto.const_mul _ $ tendsto_zero_iff_norm_tendsto_zero.mpr _), + simp_rw norm_cexp_neg_mul_sq b, + exact tendsto_exp_at_bot.comp + (tendsto.neg_const_mul_at_top (neg_lt_zero.2 hb) (tendsto_pow_at_top two_ne_zero)) }, simpa using L, end -theorem integral_gaussian (b : ℝ) : ∫ x, exp (-b * x^2) = sqrt (π / b) := +/-- The *square* of the Gaussian integral `∫ x:ℝ, exp (-b * x^2)` is equal to `π / b`. -/ +lemma integral_gaussian_sq_complex {b : ℂ} (hb : 0 < b.re) : + (∫ x:ℝ, cexp (-b * x^2)) ^ 2 = π / b := begin - /- First we deal with the crazy case where `b ≤ 0`: then both sides vanish. -/ - rcases le_or_lt b 0 with hb|hb, - { rw [integral_undef, sqrt_eq_zero_of_nonpos], - { exact div_nonpos_of_nonneg_of_nonpos pi_pos.le hb }, - { simpa only [not_lt, integrable_exp_neg_mul_sq_iff] using hb } }, - /- Assume now `b > 0`. We will show that the squares of the sides coincide. -/ - refine (sq_eq_sq _ (sqrt_nonneg _)).1 _, - { exact integral_nonneg (λ x, (exp_pos _).le) }, - /- We compute `(∫ exp(-b x^2))^2` as an integral over `ℝ^2`, and then make a polar change of - coordinates. We are left with `∫ r * exp (-b r^2)`, which has been computed in - `integral_mul_exp_neg_mul_sq` using the fact that this function has an obvious primitive. -/ + /- We compute `(∫ exp (-b x^2))^2` as an integral over `ℝ^2`, and then make a polar change + of coordinates. We are left with `∫ r * exp (-b r^2)`, which has been computed in + `integral_mul_cexp_neg_mul_sq` using the fact that this function has an obvious primitive. -/ calc - (∫ x, real.exp (-b * x^2)) ^ 2 - = ∫ p : ℝ × ℝ, exp (-b * p.1 ^ 2) * exp (-b * p.2 ^ 2) : + (∫ x:ℝ, cexp (-b * (x:ℂ)^2)) ^ 2 + = ∫ p : ℝ × ℝ, cexp (-b * ((p.1) : ℂ) ^ 2) * cexp (-b * ((p.2) : ℂ) ^ 2) : by { rw [pow_two, ← integral_prod_mul], refl } - ... = ∫ p : ℝ × ℝ, real.exp (- b * (p.1 ^ 2 + p.2^2)) : - by { congr, ext p, simp only [← real.exp_add, neg_add_rev, real.exp_eq_exp], ring } - ... = ∫ p in polar_coord.target, p.1 * exp (- b * ((p.1 * cos p.2) ^ 2 + (p.1 * sin p.2)^2)) : - (integral_comp_polar_coord_symm (λ p, exp (- b * (p.1^2 + p.2^2)))).symm - ... = (∫ r in Ioi (0 : ℝ), r * exp (-b * r^2)) * (∫ θ in Ioo (-π) π, 1) : + ... = ∫ p : ℝ × ℝ, cexp (- b * (p.1 ^ 2 + p.2 ^ 2)) : + by { congr, ext1 p, rw [← complex.exp_add, mul_add], } + ... = ∫ p in polar_coord.target, (p.1) • cexp (- b * ((p.1 * cos p.2) ^ 2 + (p.1 * sin p.2)^2)) : + begin + rw ← integral_comp_polar_coord_symm, + simp only [polar_coord_symm_apply, of_real_mul, of_real_cos, of_real_sin], + end + ... = (∫ r in Ioi (0 : ℝ), r * cexp (-b * r^2)) * (∫ θ in Ioo (-π) π, 1) : begin rw ← set_integral_prod_mul, - congr' with p, + congr' with p : 1, rw mul_one, congr, - conv_rhs { rw [← one_mul (p.1^2), ← sin_sq_add_cos_sq p.2], }, + conv_rhs { rw [← one_mul ((p.1 : ℂ)^2), ← sin_sq_add_cos_sq (p.2 : ℂ)], }, ring_exp, end - ... = π / b : + ... = ↑π / b : begin have : 0 ≤ π + π, by linarith [real.pi_pos], - simp only [integral_const, measure.restrict_apply', measurable_set_Ioo, univ_inter, this, - sub_neg_eq_add, algebra.id.smul_eq_mul, mul_one, volume_Ioo, two_mul, - ennreal.to_real_of_real, integral_mul_exp_neg_mul_sq hb, one_mul], - field_simp [hb.ne'], + simp only [integral_const, measure.restrict_apply', measurable_set_Ioo, univ_inter, + volume_Ioo, sub_neg_eq_add, ennreal.to_real_of_real, this], + rw [←two_mul, real_smul, mul_one, of_real_mul, of_real_bit0, of_real_one, + integral_mul_cexp_neg_mul_sq hb], + field_simp [(by { contrapose! hb, rw [hb, zero_re] } : b ≠ 0)], ring, end - ... = (sqrt (π / b)) ^ 2 : - by { rw sq_sqrt, exact div_nonneg pi_pos.le hb.le } end -open_locale interval - -/- The Gaussian integral on the half-line, `∫ x in Ioi 0, exp (-b * x^2)`. -/ -lemma integral_gaussian_Ioi (b : ℝ) : ∫ x in Ioi 0, exp (-b * x^2) = sqrt (π / b) / 2 := +theorem integral_gaussian (b : ℝ) : ∫ x, exp (-b * x^2) = sqrt (π / b) := begin + /- First we deal with the crazy case where `b ≤ 0`: then both sides vanish. -/ rcases le_or_lt b 0 with hb|hb, - { rw [integral_undef, sqrt_eq_zero_of_nonpos, zero_div], - exact div_nonpos_of_nonneg_of_nonpos pi_pos.le hb, - rwa [←integrable_on, integrable_on_Ioi_exp_neg_mul_sq_iff, not_lt] }, - have full_integral := integral_gaussian b, + { rw [integral_undef, sqrt_eq_zero_of_nonpos], + { exact div_nonpos_of_nonneg_of_nonpos pi_pos.le hb }, + { simpa only [not_lt, integrable_exp_neg_mul_sq_iff] using hb } }, + /- Assume now `b > 0`. Then both sides are non-negative and their squares agree. -/ + refine (sq_eq_sq _ (sqrt_nonneg _)).1 _, + { exact integral_nonneg (λ x, (exp_pos _).le) }, + rw [←of_real_inj, of_real_pow, ←integral_of_real, sq_sqrt (div_pos pi_pos hb).le, of_real_div], + convert integral_gaussian_sq_complex (by rwa of_real_re : 0 < (b:ℂ).re), + ext1 x, + rw [of_real_exp, of_real_mul, of_real_pow, of_real_neg], +end + +lemma continuous_at_gaussian_integral (b : ℂ) (hb : 0 < re b) : + continuous_at (λ c:ℂ, ∫ x:ℝ, cexp (-c * x^2)) b := +begin + let f : ℂ → ℝ → ℂ := λ (c : ℂ) (x : ℝ), cexp (-c * x ^ 2), + obtain ⟨d, hd, hd'⟩ := exists_between hb, + have f_meas : ∀ (c:ℂ), ae_strongly_measurable (f c) volume := λ c, by + { apply continuous.ae_strongly_measurable, + exact complex.continuous_exp.comp (continuous_const.mul (continuous_of_real.pow 2)) }, + have f_int : integrable (f b) volume, + { simp_rw [←integrable_norm_iff (f_meas b), norm_cexp_neg_mul_sq b], + exact integrable_exp_neg_mul_sq hb, }, + have f_cts : ∀ (x : ℝ), continuous_at (λ c, f c x) b := + λ x, (complex.continuous_exp.comp (continuous_id'.neg.mul continuous_const)).continuous_at, + have f_le_bd : ∀ᶠ (c : ℂ) in 𝓝 b, ∀ᵐ (x : ℝ), ‖f c x‖ ≤ exp (-d * x ^ 2), + { refine eventually_of_mem ((continuous_re.is_open_preimage _ is_open_Ioi).mem_nhds hd') _, + refine λ c hc, ae_of_all _ (λ x, _), + rw [norm_cexp_neg_mul_sq, exp_le_exp], + exact mul_le_mul_of_nonneg_right (neg_le_neg (le_of_lt hc)) (sq_nonneg _) }, + exact continuous_at_of_dominated (eventually_of_forall f_meas) f_le_bd + (integrable_exp_neg_mul_sq hd) (ae_of_all _ f_cts), +end + +theorem integral_gaussian_complex {b : ℂ} (hb : 0 < re b) : + ∫ x:ℝ, cexp (-b * x^2) = (π / b) ^ (1 / 2 : ℂ) := +begin + have nv : ∀ {b : ℂ}, (0 < re b) → (b ≠ 0), + { intros b hb, contrapose! hb, rw hb, simp }, + refine (convex_halfspace_re_gt 0).is_preconnected.eq_of_sq_eq + _ _ (λ c hc, _) (λ c hc, _) (by simp : 0 < re (1 : ℂ)) _ hb, + { -- integral is continuous + exact continuous_at.continuous_on continuous_at_gaussian_integral, }, + { -- `(π / b) ^ (1 / 2 : ℂ)` is continuous + refine continuous_at.continuous_on (λ b hb, (continuous_at_cpow_const (or.inl _)).comp + (continuous_at_const.div continuous_at_id (nv hb))), + rw [div_re, of_real_im, of_real_re, zero_mul, zero_div, add_zero], + exact div_pos (mul_pos pi_pos hb) (norm_sq_pos.mpr (nv hb)), }, + { -- squares of both sides agree + dsimp only [pi.pow_apply], + rw [integral_gaussian_sq_complex hc, sq], + conv_lhs { rw ←cpow_one (↑π / c)}, + rw ← cpow_add _ _ (div_ne_zero (of_real_ne_zero.mpr pi_ne_zero) (nv hc)), + norm_num }, + { -- RHS doesn't vanish + rw [ne.def, cpow_eq_zero_iff, not_and_distrib], + exact or.inl (div_ne_zero (of_real_ne_zero.mpr pi_ne_zero) (nv hc)) }, + { -- equality at 1 + have : ∀ (x : ℝ), cexp (-1 * x ^ 2) = exp (-1 * x ^ 2), + { intro x, + simp only [of_real_exp, neg_mul, one_mul, of_real_neg, of_real_pow] }, + simp_rw [this, integral_of_real], + conv_rhs { congr, rw [←of_real_one, ←of_real_div], skip, + rw [←of_real_one, ←of_real_bit0, ←of_real_div] }, + rw [←of_real_cpow, of_real_inj], + convert integral_gaussian (1 : ℝ), + { rwa [sqrt_eq_rpow] }, + { rw [div_one], exact pi_pos.le } }, +end + +/- The Gaussian integral on the half-line, `∫ x in Ioi 0, exp (-b * x^2)`, for complex `b`. -/ +lemma integral_gaussian_complex_Ioi {b : ℂ} (hb : 0 < re b) : + ∫ x:ℝ in Ioi 0, cexp (-b * x^2) = (π / b) ^ (1 / 2 : ℂ) / 2 := +begin + have full_integral := integral_gaussian_complex hb, have : measurable_set (Ioi (0:ℝ)) := measurable_set_Ioi, - rw [←integral_add_compl this (integrable_exp_neg_mul_sq hb), compl_Ioi] at full_integral, - suffices : ∫ x in Iic 0, exp (-b * x^2) = ∫ x in Ioi 0, exp (-b * x^2), + rw [←integral_add_compl this (integrable_cexp_neg_mul_sq hb), compl_Ioi] at full_integral, + suffices : ∫ x:ℝ in Iic 0, cexp (-b * x^2) = ∫ x:ℝ in Ioi 0, cexp (-b * x^2), { rw [this, ←mul_two] at full_integral, rwa eq_div_iff, exact two_ne_zero }, - have : ∀ (c : ℝ), ∫ x in 0 .. c, exp (-b * x^2) = ∫ x in -c .. 0, exp (-b * x^2), + have : ∀ (c : ℝ), ∫ x in 0 .. c, cexp (-b * x^2) = ∫ x in -c .. 0, cexp (-b * x^2), { intro c, - have := @interval_integral.integral_comp_sub_left _ _ _ _ 0 c (λ x, exp(-b * x^2)) 0, + have := @interval_integral.integral_comp_sub_left _ _ _ _ 0 c (λ x, cexp (-b * x^2)) 0, simpa [zero_sub, neg_sq, neg_zero] using this }, have t1 := interval_integral_tendsto_integral_Ioi _ - ((integrable_exp_neg_mul_sq hb).integrable_on) tendsto_id, - have t2 : tendsto (λ c:ℝ, ∫ x in 0 .. c, exp (-b * x^2)) at_top (𝓝 ∫ x in Iic 0, exp (-b * x^2)), + ((integrable_cexp_neg_mul_sq hb).integrable_on) tendsto_id, + have t2 : tendsto (λ c:ℝ, ∫ x:ℝ in 0..c, + cexp (-b * x^2)) at_top (𝓝 ∫ x:ℝ in Iic 0, cexp (-b * x^2)), { simp_rw this, refine interval_integral_tendsto_integral_Iic _ _ tendsto_neg_at_top_at_bot, - apply (integrable_exp_neg_mul_sq hb).integrable_on }, + apply (integrable_cexp_neg_mul_sq hb).integrable_on }, exact tendsto_nhds_unique t2 t1, end +/- The Gaussian integral on the half-line, `∫ x in Ioi 0, exp (-b * x^2)`, for real `b`. -/ +lemma integral_gaussian_Ioi (b : ℝ) : ∫ x in Ioi 0, exp (-b * x^2) = sqrt (π / b) / 2 := +begin + rcases le_or_lt b 0 with hb|hb, + { rw [integral_undef, sqrt_eq_zero_of_nonpos, zero_div], + exact div_nonpos_of_nonneg_of_nonpos pi_pos.le hb, + rwa [←integrable_on, integrable_on_Ioi_exp_neg_mul_sq_iff, not_lt] }, + rw [←of_real_inj, ←integral_of_real], + convert integral_gaussian_complex_Ioi (by rwa of_real_re : 0 < (b:ℂ).re), + { ext1 x, simp, }, + { rw [sqrt_eq_rpow, ←of_real_div, of_real_div, of_real_cpow], + norm_num, + exact (div_pos pi_pos hb).le, } +end + namespace complex /-- The special-value formula `Γ(1/2) = √π`, which is equivalent to the Gaussian integral. -/ @@ -210,7 +318,7 @@ begin have hh : (1 / 2 : ℂ) = ↑(1 / 2 : ℝ), { simp only [one_div, of_real_inv, of_real_bit0, of_real_one] }, have hh2 : (1 / 2 : ℂ).re = 1 / 2, - { convert complex.of_real_re (1 / 2 : ℝ) }, + { convert of_real_re (1 / 2 : ℝ) }, replace hh2 : 0 < (1 / 2 : ℂ).re := by { rw hh2, exact one_half_pos, }, rw [Gamma_eq_integral _ hh2, hh, Gamma_integral_of_real, of_real_inj, real.Gamma_integral], -- now do change-of-variables diff --git a/src/topology/algebra/field.lean b/src/topology/algebra/field.lean index 3ce7acb675118..657ca380445ed 100644 --- a/src/topology/algebra/field.lean +++ b/src/topology/algebra/field.lean @@ -185,3 +185,72 @@ lemma is_local_min.inv {f : α → β} {a : α} (h1 : is_local_min f a) (h2 : by filter_upwards [h1, h2] with z h3 h4 using (inv_le_inv h4 h2.self_of_nhds).mpr h3 end local_extr + +section preconnected +/-! Some results about functions on preconnected sets valued in a ring or field with a topology. -/ + +open set +variables {α 𝕜 : Type*} {f g : α → 𝕜} {S : set α} + [topological_space α] [topological_space 𝕜] [t1_space 𝕜] + +/-- If `f` is a function `α → 𝕜` which is continuous on a preconnected set `S`, and +`f ^ 2 = 1` on `S`, then either `f = 1` on `S`, or `f = -1` on `S`. -/ +lemma is_preconnected.eq_one_or_eq_neg_one_of_sq_eq [ring 𝕜] [no_zero_divisors 𝕜] + (hS : is_preconnected S) (hf : continuous_on f S) (hsq : eq_on (f ^ 2) 1 S) : + (eq_on f 1 S) ∨ (eq_on f (-1) S) := +begin + simp_rw [eq_on, pi.one_apply, pi.pow_apply, sq_eq_one_iff] at hsq, + -- First deal with crazy case where `S` is empty. + by_cases hSe : ∀ (x:α), x ∉ S, + { left, intros x hx, + exfalso, exact hSe x hx, }, + push_neg at hSe, + choose y hy using hSe, + suffices : ∀ (x:α), x ∈ S → f x = f y, + { rcases (hsq hy), + { left, intros z hz, rw [pi.one_apply z, ←h], exact this z hz, }, + { right, intros z hz, rw [pi.neg_apply, pi.one_apply, ←h], exact this z hz, } }, + refine λ x hx, hS.constant_of_maps_to hf (λ z hz, _) hx hy, + show f z ∈ ({-1, 1} : set 𝕜), + { exact mem_insert_iff.mpr (hsq hz).symm, }, + exact discrete_of_t1_of_finite, +end + +/-- If `f, g` are functions `α → 𝕜`, both continuous on a preconnected set `S`, with +`f ^ 2 = g ^ 2` on `S`, and `g z ≠ 0` all `z ∈ S`, then either `f = g` or `f = -g` on +`S`. -/ +lemma is_preconnected.eq_or_eq_neg_of_sq_eq [field 𝕜] [has_continuous_inv₀ 𝕜] [has_continuous_mul 𝕜] + (hS : is_preconnected S) (hf : continuous_on f S) (hg : continuous_on g S) + (hsq : eq_on (f ^ 2) (g ^ 2) S) (hg_ne : ∀ {x:α}, x ∈ S → g x ≠ 0) : + (eq_on f g S) ∨ (eq_on f (-g) S) := +begin + rcases hS.eq_one_or_eq_neg_one_of_sq_eq (hf.div hg (λ z hz, hg_ne hz)) (λ x hx, _) with h | h, + { refine or.inl (λ x hx, _), + rw ←div_eq_one_iff_eq (hg_ne hx), + exact h hx }, + { refine or.inr (λ x hx, _), + specialize h hx, + rwa [pi.div_apply, pi.neg_apply, pi.one_apply, div_eq_iff (hg_ne hx), neg_one_mul] at h, }, + { rw [pi.one_apply, div_pow, pi.div_apply, hsq hx, div_self], + exact pow_ne_zero _ (hg_ne hx) }, +end + +/-- If `f, g` are functions `α → 𝕜`, both continuous on a preconnected set `S`, with +`f ^ 2 = g ^ 2` on `S`, and `g z ≠ 0` all `z ∈ S`, then as soon as `f = g` holds at +one point of `S` it holds for all points. -/ +lemma is_preconnected.eq_of_sq_eq [field 𝕜] [has_continuous_inv₀ 𝕜] [has_continuous_mul 𝕜] + (hS : is_preconnected S) (hf : continuous_on f S) (hg : continuous_on g S) + (hsq : eq_on (f ^ 2) (g ^ 2) S) (hg_ne : ∀ {x:α}, x ∈ S → g x ≠ 0) + {y : α} (hy : y ∈ S) (hy' : f y = g y) : eq_on f g S := +λ x hx, begin + rcases hS.eq_or_eq_neg_of_sq_eq hf hg @hsq @hg_ne with h | h, + { exact h hx }, + { rw [h hy, eq_comm, ←sub_eq_zero, sub_eq_add_neg, pi.neg_apply, + neg_neg, ←mul_two, mul_eq_zero] at hy', + cases hy', -- need to handle case of `char 𝕜 = 2` separately + { exfalso, exact hg_ne hy hy' }, + { rw [h hx, pi.neg_apply, eq_comm, ←sub_eq_zero, sub_eq_add_neg, neg_neg, + ←mul_two, hy', mul_zero], } }, +end + +end preconnected diff --git a/src/topology/connected.lean b/src/topology/connected.lean index 8f43a92945c4c..19602fb7b606b 100644 --- a/src/topology/connected.lean +++ b/src/topology/connected.lean @@ -1583,3 +1583,17 @@ lemma preconnected_space_of_forall_constant (hs : ∀ f : α → bool, continuou preconnected_space α := ⟨is_preconnected_of_forall_constant (λ f hf x hx y hy, hs f (continuous_iff_continuous_on_univ.mpr hf) x y)⟩ + +/-- Refinement of `is_preconnected.constant` only assuming the map factors through a +discrete subset of the target. -/ +lemma is_preconnected.constant_of_maps_to [topological_space β] + {S : set α} (hS : is_preconnected S) {T : set β} [discrete_topology T] {f : α → β} + (hc : continuous_on f S) (hTm : maps_to f S T) + {x y : α} (hx : x ∈ S) (hy : y ∈ S) : f x = f y := +begin + let F : S → T := (λ x:S, ⟨f x.val, hTm x.property⟩), + suffices : F ⟨x, hx⟩ = F ⟨y, hy⟩, + { rw ←subtype.coe_inj at this, exact this }, + exact (is_preconnected_iff_preconnected_space.mp hS).constant + (continuous_induced_rng.mpr $ continuous_on_iff_continuous_restrict.mp hc) +end From 657df4339ae6ceada048c8a2980fb10e393143ec Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 11 Jan 2023 21:36:32 +0000 Subject: [PATCH 0214/1291] =?UTF-8?q?refactor(linear=5Falgebra/tensor=5Fpr?= =?UTF-8?q?oduct):=20make=20`lift=20f=20(x=20=E2=8A=97=E2=82=9C=20y)=20=3D?= =?UTF-8?q?=20f=20x=20y`=20true=20by=20`rfl`=20(#18121)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Since this is essentially the "primitive" recursor, it is very convenient for it to expand definitionally. With this change, the following lemmas are now rfl: * `algebra.mul'_apply` * `free_monoid.lift_comp_of` * `free_monoid.lift_eval_of` * `tensor_product.lie_module.lift_apply` * `alternating_map.dom_coprod'_apply` * `contract_left_apply` * `contract_right_apply` * `dual_tensor_hom_apply` * `derivation.tensor_product_to_tmul` * `poly_equiv_tensor.to_fun_linear_tmul_apply` * `tensor_product.lift.tmul` * `tensor_product.lift.tmul'` * `algebra.tensor_product.lift_tmul` * `algebra.tensor_product.lmul'_apply_tmul` * `tensor_product.algebra.smul_def` And one lemma is no longer rfl * `free_monoid.lift_apply` --- src/algebra/algebra/bilinear.lean | 3 +- src/algebra/free_monoid/basic.lean | 30 ++++++++++++++----- src/algebra/lie/tensor_product.lean | 2 +- .../monoidal/internal/Module.lean | 16 ++++------ src/group_theory/submonoid/membership.lean | 4 +-- src/linear_algebra/alternating.lean | 2 +- src/linear_algebra/contraction.lean | 7 +++-- src/linear_algebra/tensor_product.lean | 14 ++++----- src/ring_theory/derivation.lean | 4 +-- src/ring_theory/polynomial_algebra.lean | 7 ++--- src/ring_theory/tensor_product.lean | 13 ++++---- 11 files changed, 52 insertions(+), 50 deletions(-) diff --git a/src/algebra/algebra/bilinear.lean b/src/algebra/algebra/bilinear.lean index 3e3a3fe292da4..9a89ab27b3907 100644 --- a/src/algebra/algebra/bilinear.lean +++ b/src/algebra/algebra/bilinear.lean @@ -59,8 +59,7 @@ variables {R} @[simp] lemma mul_right_apply (a b : A) : mul_right R a b = b * a := rfl @[simp] lemma mul_left_right_apply (a b x : A) : mul_left_right R (a, b) x = a * x * b := rfl -@[simp] lemma mul'_apply {a b : A} : mul' R A (a ⊗ₜ b) = a * b := -by simp only [linear_map.mul', tensor_product.lift.tmul, mul_apply'] +@[simp] lemma mul'_apply {a b : A} : mul' R A (a ⊗ₜ b) = a * b := rfl @[simp] lemma mul_left_zero_eq_zero : mul_left R (0 : A) = 0 := diff --git a/src/algebra/free_monoid/basic.lean b/src/algebra/free_monoid/basic.lean index 7f3a305ab6003..7ccb6fdeddc5f 100644 --- a/src/algebra/free_monoid/basic.lean +++ b/src/algebra/free_monoid/basic.lean @@ -136,27 +136,41 @@ lemma hom_eq ⦃f g : free_monoid α →* M⦄ (h : ∀ x, f (of x) = g (of x)) monoid_hom.ext $ λ l, rec_on l (f.map_one.trans g.map_one.symm) $ λ x xs hxs, by simp only [h, hxs, monoid_hom.map_mul] +/-- A variant of `list.prod` that has `[x].prod = x` true definitionally. + +The purpose is to make `free_monoid.lift_eval_of` true by `rfl`. -/ +@[to_additive "A variant of `list.sum` that has `[x].sum = x` true definitionally. + +The purpose is to make `free_add_monoid.lift_eval_of` true by `rfl`."] +def prod_aux {M} [monoid M] (l : list M) : M := +l.rec_on 1 (λ x xs (_ : M), list.foldl (*) x xs) + +@[to_additive] +lemma prod_aux_eq : ∀ l : list M, free_monoid.prod_aux l = l.prod +| [] := rfl +| (x :: xs) := congr_arg (λ x, list.foldl (*) x xs) (one_mul _).symm + /-- Equivalence between maps `α → M` and monoid homomorphisms `free_monoid α →* M`. -/ @[to_additive "Equivalence between maps `α → A` and additive monoid homomorphisms `free_add_monoid α →+ A`."] def lift : (α → M) ≃ (free_monoid α →* M) := -{ to_fun := λ f, ⟨λ l, (l.to_list.map f).prod, rfl, - λ l₁ l₂, by simp only [to_list_mul, list.map_append, list.prod_append]⟩, +{ to_fun := λ f, ⟨λ l, free_monoid.prod_aux (l.to_list.map f), rfl, + λ l₁ l₂, by simp only [prod_aux_eq, to_list_mul, list.map_append, list.prod_append]⟩, inv_fun := λ f x, f (of x), - left_inv := λ f, funext $ λ x, one_mul (f x), - right_inv := λ f, hom_eq $ λ x, one_mul (f (of x)) } + left_inv := λ f, rfl, + right_inv := λ f, hom_eq $ λ x, rfl } @[simp, to_additive] lemma lift_symm_apply (f : free_monoid α →* M) : lift.symm f = f ∘ of := rfl @[to_additive] -lemma lift_apply (f : α → M) (l : free_monoid α) : lift f l = (l.to_list.map f).prod := rfl +lemma lift_apply (f : α → M) (l : free_monoid α) : lift f l = (l.to_list.map f).prod := +prod_aux_eq _ -@[to_additive] lemma lift_comp_of (f : α → M) : lift f ∘ of = f := lift.symm_apply_apply f +@[to_additive] lemma lift_comp_of (f : α → M) : lift f ∘ of = f := rfl @[simp, to_additive] -lemma lift_eval_of (f : α → M) (x : α) : lift f (of x) = f x := -congr_fun (lift_comp_of f) x +lemma lift_eval_of (f : α → M) (x : α) : lift f (of x) = f x := rfl @[simp, to_additive] lemma lift_restrict (f : free_monoid α →* M) : lift (f ∘ of) = f := diff --git a/src/algebra/lie/tensor_product.lean b/src/algebra/lie/tensor_product.lean index 2880b7c9aba29..ed0389173dcbc 100644 --- a/src/algebra/lie/tensor_product.lean +++ b/src/algebra/lie/tensor_product.lean @@ -87,7 +87,7 @@ def lift : (M →ₗ[R] N →ₗ[R] P) ≃ₗ⁅R,L⁆ (M ⊗[R] N →ₗ[R] P) @[simp] lemma lift_apply (f : M →ₗ[R] N →ₗ[R] P) (m : M) (n : N) : lift R L M N P f (m ⊗ₜ n) = f m n := -lift.equiv_apply R M N P f m n +rfl /-- A weaker form of the universal property for tensor product of modules of a Lie algebra. diff --git a/src/category_theory/monoidal/internal/Module.lean b/src/category_theory/monoidal/internal/Module.lean index e71d858040d5e..ba73dec00fd2a 100644 --- a/src/category_theory/monoidal/internal/Module.lean +++ b/src/category_theory/monoidal/internal/Module.lean @@ -66,7 +66,7 @@ instance (A : Mon_ (Module.{u} R)) : algebra R A.X := have h₂ := linear_map.congr_fun A.mul_one (a ⊗ₜ r), exact h₁.trans h₂.symm, end, - smul_def' := λ r a, by { convert (linear_map.congr_fun A.one_mul (r ⊗ₜ a)).symm, simp, }, + smul_def' := λ r a, (linear_map.congr_fun A.one_mul (r ⊗ₜ a)).symm, ..A.one } @[simp] lemma algebra_map (A : Mon_ (Module.{u} R)) (r : R) : algebra_map R A.X r = A.one r := rfl @@ -127,10 +127,8 @@ def inverse : Algebra.{u} R ⥤ Mon_ (Module.{u} R) := { obj := inverse_obj, map := λ A B f, { hom := f.to_linear_map, - one_hom' := - by { ext, dsimp, simp only [ring_hom.map_one, alg_hom.map_one] }, - mul_hom' := - by { ext, dsimp, simp only [linear_map.mul'_apply, ring_hom.map_mul, alg_hom.map_mul] } } }. + one_hom' := linear_map.ext f.commutes, + mul_hom' := tensor_product.ext $ linear_map.ext₂ $ map_mul f, } } end Mon_Module_equivalence_Algebra @@ -146,11 +144,9 @@ def Mon_Module_equivalence_Algebra : Mon_ (Module.{u} R) ≌ Algebra R := unit_iso := nat_iso.of_components (λ A, { hom := { hom := { to_fun := id, map_add' := λ x y, rfl, map_smul' := λ r a, rfl, }, - mul_hom' := by { ext, dsimp at *, - simp only [linear_map.mul'_apply, Mon_.X.ring_mul] } }, + mul_hom' := by { ext, dsimp at *, refl } }, inv := { hom := { to_fun := id, map_add' := λ x y, rfl, map_smul' := λ r a, rfl, }, - mul_hom' := by { ext, dsimp at *, - simp only [linear_map.mul'_apply, Mon_.X.ring_mul]} } }) + mul_hom' := by { ext, dsimp at *, refl } } }) (by tidy), counit_iso := nat_iso.of_components (λ A, { hom := @@ -158,7 +154,7 @@ def Mon_Module_equivalence_Algebra : Mon_ (Module.{u} R) ≌ Algebra R := map_zero' := rfl, map_add' := λ x y, rfl, map_one' := (algebra_map R A).map_one, - map_mul' := λ x y, linear_map.mul'_apply, + map_mul' := λ x y, (@linear_map.mul'_apply R _ _ _ _ _ _ x y), commutes' := λ r, rfl, }, inv := { to_fun := id, diff --git a/src/group_theory/submonoid/membership.lean b/src/group_theory/submonoid/membership.lean index 61336ecb40075..355e74a4360d6 100644 --- a/src/group_theory/submonoid/membership.lean +++ b/src/group_theory/submonoid/membership.lean @@ -291,8 +291,8 @@ by rw [free_monoid.mrange_lift, subtype.range_coe] @[to_additive] lemma closure_eq_image_prod (s : set M) : (closure s : set M) = list.prod '' {l : list M | ∀ x ∈ l, x ∈ s} := begin - rw [closure_eq_mrange, coe_mrange, ← list.range_map_coe, ← set.range_comp], - refl + rw [closure_eq_mrange, coe_mrange, ← list.range_map_coe, ← set.range_comp, function.comp], + exact congr_arg _ (funext $ free_monoid.lift_apply _), end @[to_additive] diff --git a/src/linear_algebra/alternating.lean b/src/linear_algebra/alternating.lean index d720e6127d3e7..dbbcf8a82b4bb 100644 --- a/src/linear_algebra/alternating.lean +++ b/src/linear_algebra/alternating.lean @@ -862,7 +862,7 @@ tensor_product.lift $ by lemma dom_coprod'_apply (a : alternating_map R' Mᵢ N₁ ιa) (b : alternating_map R' Mᵢ N₂ ιb) : dom_coprod' (a ⊗ₜ[R'] b) = dom_coprod a b := -by simp only [dom_coprod', tensor_product.lift.tmul, linear_map.mk₂_apply] +rfl end alternating_map diff --git a/src/linear_algebra/contraction.lean b/src/linear_algebra/contraction.lean index 29aa045beee48..13f941f92f2f8 100644 --- a/src/linear_algebra/contraction.lean +++ b/src/linear_algebra/contraction.lean @@ -48,14 +48,14 @@ def dual_tensor_hom : (module.dual R M) ⊗ N →ₗ[R] M →ₗ[R] N := variables {R M N P Q} @[simp] lemma contract_left_apply (f : module.dual R M) (m : M) : - contract_left R M (f ⊗ₜ m) = f m := by apply uncurry_apply + contract_left R M (f ⊗ₜ m) = f m := rfl @[simp] lemma contract_right_apply (f : module.dual R M) (m : M) : - contract_right R M (m ⊗ₜ f) = f m := by apply uncurry_apply + contract_right R M (m ⊗ₜ f) = f m := rfl @[simp] lemma dual_tensor_hom_apply (f : module.dual R M) (m : M) (n : N) : dual_tensor_hom R M N (f ⊗ₜ n) m = (f m) • n := -by { dunfold dual_tensor_hom, rw uncurry_apply, refl, } +rfl @[simp] lemma transpose_dual_tensor_hom (f : module.dual R M) (m : M) : dual.transpose (dual_tensor_hom R M M (f ⊗ₜ m)) = dual_tensor_hom R _ _ (dual.eval R M m ⊗ₜ f) := @@ -199,6 +199,7 @@ begin have h : function.surjective e.to_linear_map := e.surjective, refine (cancel_right h).1 _, ext p f q m, + dsimp [ltensor_hom_equiv_hom_ltensor], simp only [ltensor_hom_equiv_hom_ltensor, dual_tensor_hom_equiv, compr₂_apply, mk_apply, coe_comp, linear_equiv.coe_to_linear_map, function.comp_app, map_tmul, linear_equiv.coe_coe, dual_tensor_hom_equiv_of_basis_apply, linear_equiv.trans_apply, congr_tmul, diff --git a/src/linear_algebra/tensor_product.lean b/src/linear_algebra/tensor_product.lean index 49321d45a4bcd..f0a423cee2d83 100644 --- a/src/linear_algebra/tensor_product.lean +++ b/src/linear_algebra/tensor_product.lean @@ -403,8 +403,7 @@ add_con.add_con_gen_le $ λ x y hxy, match x, y, hxy with by simp_rw [add_monoid_hom.map_add, add_comm] end -lemma lift_aux_tmul (m n) : lift_aux f (m ⊗ₜ n) = f m n := -zero_add _ +lemma lift_aux_tmul (m n) : lift_aux f (m ⊗ₜ n) = f m n := rfl variable {f} @@ -422,11 +421,8 @@ def lift : M ⊗ N →ₗ[R] P := .. lift_aux f } variable {f} -@[simp] lemma lift.tmul (x y) : lift f (x ⊗ₜ y) = f x y := -zero_add _ - -@[simp] lemma lift.tmul' (x y) : (lift f).1 (x ⊗ₜ y) = f x y := -lift.tmul _ _ +@[simp] lemma lift.tmul (x y) : lift f (x ⊗ₜ y) = f x y := rfl +@[simp] lemma lift.tmul' (x y) : (lift f).1 (x ⊗ₜ y) = f x y := rfl theorem ext' {g h : (M ⊗[R] N) →ₗ[R] P} (H : ∀ x y, g (x ⊗ₜ y) = h (x ⊗ₜ y)) : g = h := @@ -653,11 +649,11 @@ variables [add_comm_monoid Q'] [module R Q'] lemma map_comp (f₂ : P →ₗ[R] P') (f₁ : M →ₗ[R] P) (g₂ : Q →ₗ[R] Q') (g₁ : N →ₗ[R] Q) : map (f₂.comp f₁) (g₂.comp g₁) = (map f₂ g₂).comp (map f₁ g₁) := -ext' $ λ _ _, by simp only [linear_map.comp_apply, map_tmul] +ext' $ λ _ _, rfl lemma lift_comp_map (i : P →ₗ[R] Q →ₗ[R] Q') (f : M →ₗ[R] P) (g : N →ₗ[R] Q) : (lift i).comp (map f g) = lift ((i.comp f).compl₂ g) := -ext' $ λ _ _, by simp only [lift.tmul, map_tmul, linear_map.compl₂_apply, linear_map.comp_apply] +ext' $ λ _ _, rfl local attribute [ext] ext diff --git a/src/ring_theory/derivation.lean b/src/ring_theory/derivation.lean index ae9d3f93e190b..5b4ad1574a5df 100644 --- a/src/ring_theory/derivation.lean +++ b/src/ring_theory/derivation.lean @@ -520,7 +520,7 @@ tensor_product.algebra_tensor_module.lift ((linear_map.lsmul S (S →ₗ[R] M)). lemma derivation.tensor_product_to_tmul (D : derivation R S M) (s t : S) : D.tensor_product_to (s ⊗ₜ t) = s • D t := -tensor_product.lift.tmul s t +rfl lemma derivation.tensor_product_to_mul (D : derivation R S M) (x y : S ⊗[R] S) : D.tensor_product_to (x * y) = tensor_product.lmul' R x • D.tensor_product_to y + @@ -815,7 +815,7 @@ begin { generalize : f x = y, obtain ⟨y, rfl⟩ := ideal.quotient.mk_surjective y, refl }, have e₂ : x = kaehler_differential.quotient_cotangent_ideal_ring_equiv R S (is_scalar_tower.to_alg_hom R S _ x), - { exact ((tensor_product.lmul'_apply_tmul x 1).trans (mul_one x)).symm }, + { exact (mul_one x).symm }, split, { intro e, exact (e₁.trans (@ring_equiv.congr_arg _ _ _ _ _ _ diff --git a/src/ring_theory/polynomial_algebra.lean b/src/ring_theory/polynomial_algebra.lean index 323fb6f7d9f6f..f41aa6cf97f72 100644 --- a/src/ring_theory/polynomial_algebra.lean +++ b/src/ring_theory/polynomial_algebra.lean @@ -72,7 +72,7 @@ tensor_product.lift (to_fun_bilinear R A) @[simp] lemma to_fun_linear_tmul_apply (a : A) (p : R[X]) : - to_fun_linear R A (a ⊗ₜ[R] p) = to_fun_bilinear R A a p := lift.tmul _ _ + to_fun_linear R A (a ⊗ₜ[R] p) = to_fun_bilinear R A a p := rfl -- We apparently need to provide the decidable instance here -- in order to successfully rewrite by this lemma. @@ -125,10 +125,7 @@ alg_hom_of_linear_map_tensor_product @[simp] lemma to_fun_alg_hom_apply_tmul (a : A) (p : R[X]) : to_fun_alg_hom R A (a ⊗ₜ[R] p) = p.sum (λ n r, monomial n (a * (algebra_map R A) r)) := -begin - dsimp [to_fun_alg_hom], - rw [to_fun_linear_tmul_apply, to_fun_bilinear_apply_eq_sum], -end +to_fun_bilinear_apply_eq_sum R A _ _ /-- (Implementation detail.) diff --git a/src/ring_theory/tensor_product.lean b/src/ring_theory/tensor_product.lean index 19dfcb2c94568..24e33059f459e 100644 --- a/src/ring_theory/tensor_product.lean +++ b/src/ring_theory/tensor_product.lean @@ -117,7 +117,7 @@ the given bilinear map `M →[A] N →[R] P`. -/ @[simp] lemma lift_tmul (f : M →ₗ[A] (N →ₗ[R] P)) (x : M) (y : N) : lift f (x ⊗ₜ y) = f x y := -lift.tmul' x y +rfl variables (R A M N P) /-- Heterobasic version of `tensor_product.uncurry`: @@ -780,15 +780,15 @@ variables {R} lemma lmul'_to_linear_map : (lmul' R : _ →ₐ[R] S).to_linear_map = linear_map.mul' R S := rfl -@[simp] lemma lmul'_apply_tmul (a b : S) : lmul' R (a ⊗ₜ[R] b) = a * b := linear_map.mul'_apply +@[simp] lemma lmul'_apply_tmul (a b : S) : lmul' R (a ⊗ₜ[R] b) = a * b := rfl @[simp] lemma lmul'_comp_include_left : (lmul' R : _ →ₐ[R] S).comp include_left = alg_hom.id R S := -alg_hom.ext $ λ _, (lmul'_apply_tmul _ _).trans (_root_.mul_one _) +alg_hom.ext $ _root_.mul_one @[simp] lemma lmul'_comp_include_right : (lmul' R : _ →ₐ[R] S).comp include_right = alg_hom.id R S := -alg_hom.ext $ λ _, (lmul'_apply_tmul _ _).trans (_root_.one_mul _) +alg_hom.ext $ _root_.one_mul /-- If `S` is commutative, for a pair of morphisms `f : A →ₐ[R] S`, `g : B →ₐ[R] S`, @@ -926,8 +926,7 @@ tensor_product.lift map_smul' := λ n r, by { ext, simp only [ring_hom.id_apply, linear_map.smul_apply, smul_assoc] } } lemma module_aux_apply (a : A) (b : B) (m : M) : - module_aux (a ⊗ₜ[R] b) m = a • b • m := -by simp [module_aux] + module_aux (a ⊗ₜ[R] b) m = a • b • m := rfl variables [smul_comm_class A B M] @@ -977,6 +976,6 @@ protected def module : module (A ⊗[R] B) M := local attribute [instance] tensor_product.algebra.module -lemma smul_def (a : A) (b : B) (m : M) : (a ⊗ₜ[R] b) • m = a • b • m := module_aux_apply a b m +lemma smul_def (a : A) (b : B) (m : M) : (a ⊗ₜ[R] b) • m = a • b • m := rfl end tensor_product.algebra From 36a352116860fe906561f6cc3dd2b1e68571043f Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 12 Jan 2023 00:39:42 +0000 Subject: [PATCH 0215/1291] feat(group_theory/group_action/support): Support under an action (#17055) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Given an action of a group `G` on a type `α`, we say that a set `s : set α ` **supports** an element `a : α` if, for all `g`, `g` fixes `a` if `g` fixes `s` pointwise. --- src/group_theory/group_action/support.lean | 53 ++++++++++++++++++++++ 1 file changed, 53 insertions(+) create mode 100644 src/group_theory/group_action/support.lean diff --git a/src/group_theory/group_action/support.lean b/src/group_theory/group_action/support.lean new file mode 100644 index 0000000000000..bf39f8da16b97 --- /dev/null +++ b/src/group_theory/group_action/support.lean @@ -0,0 +1,53 @@ +/- +Copyright (c) 2022 Yaël Dillies. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies +-/ +import data.set.pointwise.smul + +/-! +# Support of an element under an action action + +Given an action of a group `G` on a type `α`, we say that a set `s : set α` supports an element +`a : α` if, for all `g` that fix `s` pointwise, `g` fixes `a`. + +This is crucial in Fourier-Motzkin constructions. +-/ + +open_locale pointwise + +variables {G H α β : Type*} + +namespace mul_action +section has_smul +variables (G) [has_smul G α] [has_smul G β] + +/-- A set `s` supports `b` if `g • b = b` whenever `g • a = a` for all `a ∈ s`. -/ +@[to_additive "A set `s` supports `b` if `g +ᵥ b = b` whenever `g +ᵥ a = a` for all `a ∈ s`."] +def supports (s : set α) (b : β) := ∀ g : G, (∀ ⦃a⦄, a ∈ s → g • a = a) → g • b = b + +variables {s t : set α} {a : α} {b : β} + +@[to_additive] lemma supports_of_mem (ha : a ∈ s) : supports G s a := λ g h, h ha + +variables {G} + +@[to_additive] lemma supports.mono (h : s ⊆ t) (hs : supports G s b) : supports G t b := +λ g hg, hs _ $ λ a ha, hg $ h ha + +end has_smul + +variables [group H] [has_smul G α] [has_smul G β] [mul_action H α] [has_smul H β] + [smul_comm_class G H β] [smul_comm_class G H α] {s t : set α} {b : β} + +-- TODO: This should work without `smul_comm_class` +@[to_additive] lemma supports.smul (g : H) (h : supports G s b) : supports G (g • s) (g • b) := +begin + rintro g' hg', + rw [smul_comm, h], + rintro a ha, + have := set.ball_image_iff.1 hg' a ha, + rwa [smul_comm, smul_left_cancel_iff] at this, +end + +end mul_action From f835503164c75764d4165d9e630bf27694d73ec6 Mon Sep 17 00:00:00 2001 From: Devon Tuma Date: Thu, 12 Jan 2023 00:39:44 +0000 Subject: [PATCH 0216/1291] feat(topology/instances/ennreal): curried version of `ennreal.tsum_prod` (#18124) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Add a version of `ennreal.tsum_prod` for functions `f : α × β → ℝ≥0∞` rather than `f : α → β → ℝ≥0∞`. --- src/topology/instances/ennreal.lean | 5 ++++- 1 file changed, 4 insertions(+), 1 deletion(-) diff --git a/src/topology/instances/ennreal.lean b/src/topology/instances/ennreal.lean index 9d679d95c2d7c..8d5af6c912c16 100644 --- a/src/topology/instances/ennreal.lean +++ b/src/topology/instances/ennreal.lean @@ -759,7 +759,10 @@ protected lemma tsum_sigma' {β : α → Type*} (f : (Σ a, β a) → ℝ≥0∞ ∑'p:(Σa, β a), f p = ∑'a b, f ⟨a, b⟩ := tsum_sigma' (assume b, ennreal.summable) ennreal.summable -protected lemma tsum_prod {f : α → β → ℝ≥0∞} : ∑'p:α×β, f p.1 p.2 = ∑'a, ∑'b, f a b := +protected lemma tsum_prod {f : α → β → ℝ≥0∞} : ∑' p : α × β, f p.1 p.2 = ∑' a b, f a b := +tsum_prod' ennreal.summable $ λ _, ennreal.summable + +protected lemma tsum_prod' {f : α × β → ℝ≥0∞} : ∑' p : α × β, f p = ∑' a b, f (a, b) := tsum_prod' ennreal.summable $ λ _, ennreal.summable protected lemma tsum_comm {f : α → β → ℝ≥0∞} : ∑'a, ∑'b, f a b = ∑'b, ∑'a, f a b := From ec322deb9ba5aad978f862669053069b7957c31d Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 12 Jan 2023 02:58:09 +0000 Subject: [PATCH 0217/1291] fix(tactic/{norm_num, abel}): do not do syntactic matches on typeclass instances (#18129) Specifically, this remove syntactic matches on `has_pow` and `has_smul` instances in favor of unification. In all cases there is a cache we can exploit to avoid looking up the preferred instance from scratch every time. --- scripts/nolints.txt | 3 +++ src/tactic/abel.lean | 20 ++++++++++++++++---- src/tactic/norm_num.lean | 13 +++++++++---- src/tactic/ring_exp.lean | 13 +++++++------ test/abel.lean | 11 +++++++++++ test/norm_num.lean | 5 +++++ test/ring_exp.lean | 4 ++++ 7 files changed, 55 insertions(+), 14 deletions(-) diff --git a/scripts/nolints.txt b/scripts/nolints.txt index eb4945fe58bf8..2b9acf01a96c8 100644 --- a/scripts/nolints.txt +++ b/scripts/nolints.txt @@ -416,6 +416,9 @@ apply_nolint tactic.abel.eval_add doc_blame apply_nolint tactic.abel.eval_atom doc_blame apply_nolint tactic.abel.eval_neg doc_blame apply_nolint tactic.abel.eval_smul doc_blame +apply_nolint tactic.abel.int_smul_instg doc_blame +apply_nolint tactic.abel.nat_smul_inst doc_blame +apply_nolint tactic.abel.nat_smul_instg doc_blame apply_nolint tactic.abel.normal_expr doc_blame apply_nolint tactic.abel.normal_expr.e doc_blame apply_nolint tactic.abel.normal_expr.pp doc_blame diff --git a/src/tactic/abel.lean b/src/tactic/abel.lean index 8e97c0ab87822..a7fb931bd1e6e 100644 --- a/src/tactic/abel.lean +++ b/src/tactic/abel.lean @@ -189,6 +189,10 @@ meta def eval_neg (c : context) : normal_expr → tactic (normal_expr × expr) return (term' c (n', -n.2) x a', c.app ``term_neg c.inst [n.1, x, a, n', a', h₁, h₂]) +def nat_smul_inst {α} [add_comm_monoid α] : has_smul ℕ α := by apply_instance +def nat_smul_instg {α} [add_comm_group α] : has_smul ℕ α := by apply_instance +def int_smul_instg {α} [add_comm_group α] : has_smul ℤ α := by apply_instance + def smul {α} [add_comm_monoid α] (n : ℕ) (x : α) : α := n • x def smulg {α} [add_comm_group α] (n : ℤ) (x : α) : α := n • x @@ -309,10 +313,18 @@ meta def eval (c : context) : expr → tactic (normal_expr × expr) guardb c.is_group, (e', p) ← eval $ c.iapp ``smul [e₁, e₂], return (e', c.app ``unfold_zsmul c.inst [e₁, e₂, e', p]) -| e@`(@has_smul.smul nat _ add_monoid.has_smul_nat %%e₁ %%e₂) := - eval_smul' c eval ff e e₁ e₂ -| e@`(@has_smul.smul int _ sub_neg_monoid.has_smul_int %%e₁ %%e₂) := - eval_smul' c eval tt e e₁ e₂ +| e@`(@has_smul.smul nat %%α %%inst %%e₁ %%e₂) := do + let inst' := c.iapp ``nat_smul_inst [], + mcond (succeeds (is_def_eq inst inst')) + (eval_smul' c eval ff e e₁ e₂) + (eval_atom c e) +| e@`(@has_smul.smul int %%α %%inst %%e₁ %%e₂) := do + -- if we're not a group there's no canonical instance available + tt ← pure c.is_group | eval_atom c e, + let inst' := c.app ``int_smul_instg c.inst [], + mcond (succeeds (is_def_eq inst inst')) + (eval_smul' c eval tt e e₁ e₂) + (eval_atom c e) | e@`(smul %%e₁ %%e₂) := eval_smul' c eval ff e e₁ e₂ | e@`(smulg %%e₁ %%e₂) := eval_smul' c eval tt e e₁ e₂ | e@`(@has_zero.zero _ _) := mcond (succeeds (is_def_eq e c.α0)) diff --git a/src/tactic/norm_num.lean b/src/tactic/norm_num.lean index 270c402ad8de1..eaeb50dd01860 100644 --- a/src/tactic/norm_num.lean +++ b/src/tactic/norm_num.lean @@ -1078,12 +1078,17 @@ meta def prove_zpow (ic zc nc : instance_cache) (a : expr) (na : ℚ) (b : expr) /-- Evaluates expressions of the form `a ^ b`, `monoid.npow a b` or `nat.pow a b`. -/ meta def eval_pow : expr → tactic (expr × expr) -| `(@has_pow.pow %%α _ %%m %%e₁ %%e₂) := do +| `(@has_pow.pow %%α %%β %%m %%e₁ %%e₂) := do n₁ ← e₁.to_rat, c ← mk_instance_cache α, - match m with - | `(@monoid.has_pow %%_ %%_) := prod.snd <$> prove_pow e₁ n₁ c e₂ - | `(@div_inv_monoid.has_pow %%_ %%_) := do + match β with + | `(ℕ) := do + (c, m') ← c.mk_app ``monoid.has_pow [], + is_def_eq m m', + prod.snd <$> prove_pow e₁ n₁ c e₂ + | `(ℤ) := do + (c, m') ← c.mk_app ``div_inv_monoid.has_pow [], + is_def_eq m m', zc ← mk_instance_cache `(ℤ), nc ← mk_instance_cache `(ℕ), (prod.snd ∘ prod.snd ∘ prod.snd) <$> prove_zpow c zc nc e₁ n₁ e₂ diff --git a/src/tactic/ring_exp.lean b/src/tactic/ring_exp.lean index c4c6e06c08079..e39e2ff0db77e 100644 --- a/src/tactic/ring_exp.lean +++ b/src/tactic/ring_exp.lean @@ -1408,16 +1408,17 @@ meta def eval : expr → ring_exp_m (ex sum) pf ← mk_app_class ``div_pf dri [ps, qs, psqs.pretty, psqs_pf], pure (psqs.set_info e pf)) <|> eval_base e | e@`(@has_pow.pow _ _ %%hp_instance %%ps %%qs) := do + ctx ← get_context, ps' ← eval ps, qs' ← in_exponent $ eval qs, psqs ← pow ps' qs', psqs_pf ← psqs.proof_term, - (do has_pow_pf ← match hp_instance with - | `(monoid.has_pow) := lift $ mk_eq_refl e - | _ := lift $ fail "has_pow instance must be nat.has_pow or monoid.has_pow" - end, - pf ← lift $ mk_eq_trans has_pow_pf psqs_pf, - pure $ psqs.set_info e pf) <|> eval_base e + (do + lift (is_def_eq hp_instance ctx.info_b.hp_instance + <|> fail "has_pow instance must be nat.has_pow or monoid.has_pow"), + has_pow_pf ← lift $ mk_eq_refl e, + pf ← lift $ mk_eq_trans has_pow_pf psqs_pf, + pure $ psqs.set_info e pf) <|> eval_base e | ps := eval_base ps /-- diff --git a/test/abel.lean b/test/abel.lean index 5cb28d799f150..db60f0c98840b 100644 --- a/test/abel.lean +++ b/test/abel.lean @@ -1,4 +1,5 @@ import tactic.abel +import algebra.group.pi variables {α : Type*} {a b : α} example [add_comm_monoid α] : a + (b + a) = a + a + b := by abel @@ -11,6 +12,16 @@ example [add_comm_group α] (a : α) : 0 + a = a := by abel1 example [add_comm_group α] (n : ℕ) (a : α) : n • a = n • a := by abel1 example [add_comm_group α] (n : ℕ) (a : α) : 0 + n • a = n • a := by abel1 +-- instances do not have to syntactically be +-- `add_monoid.has_smul_nat` or `sub_neg_monoid.has_smul_int` +example [add_comm_monoid α] (x : ℕ → α) : ((2 : ℕ) • x) = x + x := by abel1 +example [add_comm_group α] (x : ℕ → α) : ((2 : ℕ) • x) = x + x := by abel1 +example [add_comm_group α] (x : ℕ → α) : ((2 : ℤ) • x) = x + x := by abel1 + +-- even if there's an instance we don't recognize, we treat it as an atom +example [add_comm_group α] [has_smul ℕ α] (x : ℕ → α) : + ((2 : ℕ) • x) + ((2 : ℕ) • x) = (2 : ℤ) • ((2 : ℕ) • x) := by abel1 + -- `abel!` should see through terms that are definitionally equal, def id' (x : α) := x example [add_comm_group α] : a + b - b - id' a = 0 := diff --git a/test/norm_num.lean b/test/norm_num.lean index 7752d113a38aa..d2b98fdec5b84 100644 --- a/test/norm_num.lean +++ b/test/norm_num.lean @@ -5,6 +5,7 @@ Authors: Simon Hudon, Mario Carneiro -/ import tactic.norm_num +import algebra.ring.pi /-! # Tests for `norm_num` extensions @@ -303,6 +304,10 @@ example : (- ((- (((66 - 86) - 36) / 94) - 3) / - - (77 / (56 - - - 79))) + 87) example : 2 ^ 13 - 1 = int.of_nat 8191 := by norm_num +-- `^` and `•` do not have to match `monoid.has_pow` and `add_monoid.has_smul` syntactically +example {α} [ring α] : (2 ^ 3 : ℕ → α) = 8 := by norm_num +example {α} [ring α] : (2 • 3 : ℕ → α) = 6 := by norm_num + /-! Test the behaviour of removing one `norm_num` extension tactic. -/ section remove_extension diff --git a/test/ring_exp.lean b/test/ring_exp.lean index 04905ce9685f8..2485f20493a06 100644 --- a/test/ring_exp.lean +++ b/test/ring_exp.lean @@ -1,6 +1,7 @@ import tactic.ring_exp import tactic.zify import algebra.group_with_zero.power +import algebra.ring.pi import tactic.field_simp universes u @@ -75,6 +76,9 @@ example (a b : ℚ) : (a * b) ^ 1000000 = (b * a) ^ 1000000 := by ring_exp example (n : ℕ) : 2 ^ (n + 1 + 1) = 2 * 2 ^ (n + 1) := by ring_exp_eq +-- power does not have to be a syntactic match to `monoid.has_pow` +example {α} [comm_ring α] (x : ℕ → α) : (x ^ 2 * x) = x ^ 3 := by ring_exp + end exponentiation section power_of_sum From 579cdfe0a0b8782a289ecda5e3b17b3773d580d5 Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Thu, 12 Jan 2023 09:11:52 +0000 Subject: [PATCH 0218/1291] refactor(linear_algebra/affine_space/*): make `affine_basis` more elementary (#18141) In the linear algebra development of mathlib, `basis` is more elementary than `finite_dimension`, and matrix results are not imported to the main theory until determinants. This PR changes the structure of the affine linear algebra development to match that. I think this is worth doing in any case, but my motivation is to decrease the length of the [current longest chain in mathlib](https://tqft.net/mathlib4/2023-01-12/all) and particularly the length of the chain to `analysis.convex.specific_functions`, which is imported by measure theory. This actually only reduces those chains in length slightly, since there is a second nearly-as-long path from `data.set.finite` to `analysis.convex.specific_functions` which passes through topology, metric spaces, and normed spaces, rather than through linear algebra, noetherian rings, free modules, and matrix theory. But that one can be studied later. --- .../normed_space/add_torsor_bases.lean | 2 +- src/linear_algebra/affine_space/basis.lean | 176 +----------------- .../affine_space/finite_dimensional.lean | 47 ++++- src/linear_algebra/affine_space/matrix.lean | 172 +++++++++++++++++ 4 files changed, 221 insertions(+), 176 deletions(-) create mode 100644 src/linear_algebra/affine_space/matrix.lean diff --git a/src/analysis/normed_space/add_torsor_bases.lean b/src/analysis/normed_space/add_torsor_bases.lean index e5b9b25a7e6a0..87ed4056ec2a8 100644 --- a/src/analysis/normed_space/add_torsor_bases.lean +++ b/src/analysis/normed_space/add_torsor_bases.lean @@ -6,7 +6,7 @@ Authors: Oliver Nash 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 /-! # Bases in normed affine spaces. diff --git a/src/linear_algebra/affine_space/basis.lean b/src/linear_algebra/affine_space/basis.lean index 307c61c787c4a..56ad63b78851c 100644 --- a/src/linear_algebra/affine_space/basis.lean +++ b/src/linear_algebra/affine_space/basis.lean @@ -4,8 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Oliver Nash -/ import linear_algebra.affine_space.independent -import linear_algebra.affine_space.finite_dimensional -import linear_algebra.determinant +import linear_algebra.basis /-! # Affine bases and barycentric coordinates @@ -39,7 +38,7 @@ barycentric coordinate of `q : P` is `1 - fᵢ (q -ᵥ p i)`. -/ -open_locale affine big_operators matrix +open_locale affine big_operators open set universes u₁ u₂ u₃ u₄ @@ -223,178 +222,19 @@ noncomputable def coords : P →ᵃ[k] ι → k := b.coords q i = b.coord i q := rfl -/-- Given an affine basis `p`, and a family of points `q : ι' → P`, this is the matrix whose -rows are the barycentric coordinates of `q` with respect to `p`. - -It is an affine equivalent of `basis.to_matrix`. -/ -noncomputable def to_matrix {ι' : Type*} (q : ι' → P) : matrix ι' ι k := -λ i j, b.coord j (q i) - -@[simp] lemma to_matrix_apply {ι' : Type*} (q : ι' → P) (i : ι') (j : ι) : - b.to_matrix q i j = b.coord j (q i) := -rfl - -@[simp] lemma to_matrix_self [decidable_eq ι] : - b.to_matrix b.points = (1 : matrix ι ι k) := -begin - ext i j, - rw [to_matrix_apply, coord_apply, matrix.one_eq_pi_single, pi.single_apply], -end - -variables {ι' : Type*} [fintype ι'] [fintype ι] (b₂ : affine_basis ι k P) - -lemma to_matrix_row_sum_one {ι' : Type*} (q : ι' → P) (i : ι') : - ∑ j, b.to_matrix q i j = 1 := -by simp - -/-- Given a family of points `p : ι' → P` and an affine basis `b`, if the matrix whose rows are the -coordinates of `p` with respect `b` has a right inverse, then `p` is affine independent. -/ -lemma affine_independent_of_to_matrix_right_inv [decidable_eq ι'] - (p : ι' → P) {A : matrix ι ι' k} (hA : (b.to_matrix p) ⬝ A = 1) : affine_independent k p := -begin - rw affine_independent_iff_eq_of_fintype_affine_combination_eq, - intros w₁ w₂ hw₁ hw₂ hweq, - have hweq' : (b.to_matrix p).vec_mul w₁ = (b.to_matrix p).vec_mul w₂, - { ext j, - change ∑ i, (w₁ i) • (b.coord j (p i)) = ∑ i, (w₂ i) • (b.coord j (p i)), - rw [← finset.univ.affine_combination_eq_linear_combination _ _ hw₁, - ← finset.univ.affine_combination_eq_linear_combination _ _ hw₂, - ← finset.univ.map_affine_combination p w₁ hw₁, - ← finset.univ.map_affine_combination p w₂ hw₂, hweq], }, - replace hweq' := congr_arg (λ w, A.vec_mul w) hweq', - simpa only [matrix.vec_mul_vec_mul, ← matrix.mul_eq_mul, hA, matrix.vec_mul_one] using hweq', -end - -/-- Given a family of points `p : ι' → P` and an affine basis `b`, if the matrix whose rows are the -coordinates of `p` with respect `b` has a left inverse, then `p` spans the entire space. -/ -lemma affine_span_eq_top_of_to_matrix_left_inv [decidable_eq ι] [nontrivial k] - (p : ι' → P) {A : matrix ι ι' k} (hA : A ⬝ b.to_matrix p = 1) : affine_span k (range p) = ⊤ := -begin - suffices : ∀ i, b.points i ∈ affine_span k (range p), - { rw [eq_top_iff, ← b.tot, affine_span_le], - rintros q ⟨i, rfl⟩, - exact this i, }, - intros i, - have hAi : ∑ j, A i j = 1, - { calc ∑ j, A i j = ∑ j, (A i j) * ∑ l, b.to_matrix p j l : by simp - ... = ∑ j, ∑ l, (A i j) * b.to_matrix p j l : by simp_rw finset.mul_sum - ... = ∑ l, ∑ j, (A i j) * b.to_matrix p j l : by rw finset.sum_comm - ... = ∑ l, (A ⬝ b.to_matrix p) i l : rfl - ... = 1 : by simp [hA, matrix.one_apply, finset.filter_eq], }, - have hbi : b.points i = finset.univ.affine_combination p (A i), - { apply b.ext_elem, - intros j, - rw [b.coord_apply, finset.univ.map_affine_combination _ _ hAi, - finset.univ.affine_combination_eq_linear_combination _ _ hAi], - change _ = (A ⬝ b.to_matrix p) i j, - simp_rw [hA, matrix.one_apply, @eq_comm _ i j] }, - rw hbi, - exact affine_combination_mem_affine_span hAi p, -end - -/-- A change of basis formula for barycentric coordinates. - -See also `affine_basis.to_matrix_inv_mul_affine_basis_to_matrix`. -/ -@[simp] lemma to_matrix_vec_mul_coords (x : P) : - (b.to_matrix b₂.points).vec_mul (b₂.coords x) = b.coords x := -begin - ext j, - change _ = b.coord j x, - conv_rhs { rw ← b₂.affine_combination_coord_eq_self x, }, - rw finset.map_affine_combination _ _ _ (b₂.sum_coord_apply_eq_one x), - simp [matrix.vec_mul, matrix.dot_product, to_matrix_apply, coords], -end - -variables [decidable_eq ι] - -lemma to_matrix_mul_to_matrix : - (b.to_matrix b₂.points) ⬝ (b₂.to_matrix b.points) = 1 := -begin - ext l m, - change (b₂.to_matrix b.points).vec_mul (b.coords (b₂.points l)) m = _, - rw [to_matrix_vec_mul_coords, coords_apply, ← to_matrix_apply, to_matrix_self], -end - -lemma is_unit_to_matrix : - is_unit (b.to_matrix b₂.points) := -⟨{ val := b.to_matrix b₂.points, - inv := b₂.to_matrix b.points, - val_inv := b.to_matrix_mul_to_matrix b₂, - inv_val := b₂.to_matrix_mul_to_matrix b, }, rfl⟩ - -lemma is_unit_to_matrix_iff [nontrivial k] (p : ι → P) : - is_unit (b.to_matrix p) ↔ affine_independent k p ∧ affine_span k (range p) = ⊤ := -begin - split, - { rintros ⟨⟨B, A, hA, hA'⟩, (rfl : B = b.to_matrix p)⟩, - rw matrix.mul_eq_mul at hA hA', - exact ⟨b.affine_independent_of_to_matrix_right_inv p hA, - b.affine_span_eq_top_of_to_matrix_left_inv p hA'⟩, }, - { rintros ⟨h_tot, h_ind⟩, - let b' : affine_basis ι k P := ⟨p, h_tot, h_ind⟩, - change is_unit (b.to_matrix b'.points), - exact b.is_unit_to_matrix b', }, -end - end ring -section comm_ring - -variables [comm_ring k] [module k V] [decidable_eq ι] [fintype ι] -variables (b b₂ : affine_basis ι k P) - -/-- A change of basis formula for barycentric coordinates. - -See also `affine_basis.to_matrix_vec_mul_coords`. -/ -@[simp] lemma to_matrix_inv_vec_mul_to_matrix (x : P) : - (b.to_matrix b₂.points)⁻¹.vec_mul (b.coords x) = b₂.coords x := -begin - have hu := b.is_unit_to_matrix b₂, - rw matrix.is_unit_iff_is_unit_det at hu, - rw [← b.to_matrix_vec_mul_coords b₂, matrix.vec_mul_vec_mul, matrix.mul_nonsing_inv _ hu, - matrix.vec_mul_one], -end - -/-- If we fix a background affine basis `b`, then for any other basis `b₂`, we can characterise -the barycentric coordinates provided by `b₂` in terms of determinants relative to `b`. -/ -lemma det_smul_coords_eq_cramer_coords (x : P) : - (b.to_matrix b₂.points).det • b₂.coords x = (b.to_matrix b₂.points)ᵀ.cramer (b.coords x) := -begin - have hu := b.is_unit_to_matrix b₂, - rw matrix.is_unit_iff_is_unit_det at hu, - rw [← b.to_matrix_inv_vec_mul_to_matrix, matrix.det_smul_inv_vec_mul_eq_cramer_transpose _ _ hu], -end - -end comm_ring - section division_ring variables [division_ring k] [module k V] include V -protected lemma finite_dimensional [finite ι] (b : affine_basis ι k P) : finite_dimensional k V := -let ⟨i⟩ := b.nonempty in finite_dimensional.of_fintype_basis (b.basis_of i) - -protected lemma finite [finite_dimensional k V] (b : affine_basis ι k P) : finite ι := -finite_of_fin_dim_affine_independent k b.ind - -protected lemma finite_set [finite_dimensional k V] {s : set ι} (b : affine_basis s k P) : - s.finite := -finite_set_of_fin_dim_affine_independent k b.ind - @[simp] lemma coord_apply_centroid [char_zero k] (b : affine_basis ι k P) {s : finset ι} {i : ι} (hi : i ∈ s) : b.coord i (s.centroid k b.points) = (s.card : k) ⁻¹ := by rw [finset.centroid, b.coord_apply_combination_of_mem hi (s.sum_centroid_weights_eq_one_of_nonempty _ ⟨i, hi⟩), finset.centroid_weights] -lemma card_eq_finrank_add_one [fintype ι] (b : affine_basis ι k P) : - fintype.card ι = finite_dimensional.finrank k V + 1 := -begin - haveI := b.finite_dimensional, - exact b.ind.affine_span_eq_top_iff_card_eq_finrank_add_one.mp b.tot -end - lemma exists_affine_subbasis {t : set P} (ht : affine_span k t = ⊤) : ∃ (s ⊆ t) (b : affine_basis ↥s k P), b.points = coe := begin @@ -408,18 +248,6 @@ variables (k V P) lemma exists_affine_basis : ∃ (s : set P) (b : affine_basis ↥s k P), b.points = coe := let ⟨s, _, hs⟩ := exists_affine_subbasis (affine_subspace.span_univ k V P) in ⟨s, hs⟩ -variables {k V P} - -lemma exists_affine_basis_of_finite_dimensional [fintype ι] [finite_dimensional k V] - (h : fintype.card ι = finite_dimensional.finrank k V + 1) : - nonempty (affine_basis ι k P) := -begin - obtain ⟨s, b, hb⟩ := affine_basis.exists_affine_basis k V P, - lift s to finset P using b.finite_set, - refine ⟨b.comp_equiv $ fintype.equiv_of_card_eq _⟩, - rw [h, ← b.card_eq_finrank_add_one] -end - 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 ca021cbd7b0e0..568388f982953 100644 --- a/src/linear_algebra/affine_space/finite_dimensional.lean +++ b/src/linear_algebra/affine_space/finite_dimensional.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 Joseph Myers. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Joseph Myers -/ -import linear_algebra.affine_space.independent +import linear_algebra.affine_space.basis import linear_algebra.finite_dimensional /-! @@ -727,3 +727,48 @@ lemma coplanar_triple (p₁ p₂ p₃ : P) : coplanar k ({p₁, p₂, p₃} : se (collinear_pair k p₂ p₃).coplanar_insert p₁ end division_ring + +namespace affine_basis + +universes u₁ u₂ u₃ u₄ + +variables {ι : Type u₁} {k : Type u₂} {V : Type u₃} {P : Type u₄} +variables [add_comm_group V] [affine_space V P] + +section division_ring + +variables [division_ring k] [module k V] +include V + +protected lemma finite_dimensional [finite ι] (b : affine_basis ι k P) : finite_dimensional k V := +let ⟨i⟩ := b.nonempty in finite_dimensional.of_fintype_basis (b.basis_of i) + +protected lemma finite [finite_dimensional k V] (b : affine_basis ι k P) : finite ι := +finite_of_fin_dim_affine_independent k b.ind + +protected lemma finite_set [finite_dimensional k V] {s : set ι} (b : affine_basis s k P) : + s.finite := +finite_set_of_fin_dim_affine_independent k b.ind + +lemma card_eq_finrank_add_one [fintype ι] (b : affine_basis ι k P) : + fintype.card ι = finite_dimensional.finrank k V + 1 := +begin + haveI := b.finite_dimensional, + exact b.ind.affine_span_eq_top_iff_card_eq_finrank_add_one.mp b.tot +end + +variables {k V P} + +lemma exists_affine_basis_of_finite_dimensional [fintype ι] [finite_dimensional k V] + (h : fintype.card ι = finite_dimensional.finrank k V + 1) : + nonempty (affine_basis ι k P) := +begin + obtain ⟨s, b, hb⟩ := affine_basis.exists_affine_basis k V P, + lift s to finset P using b.finite_set, + refine ⟨b.comp_equiv $ fintype.equiv_of_card_eq _⟩, + rw [h, ← b.card_eq_finrank_add_one] +end + +end division_ring + +end affine_basis diff --git a/src/linear_algebra/affine_space/matrix.lean b/src/linear_algebra/affine_space/matrix.lean new file mode 100644 index 0000000000000..b0109bc3f186c --- /dev/null +++ b/src/linear_algebra/affine_space/matrix.lean @@ -0,0 +1,172 @@ +/- +Copyright (c) 2021 Oliver Nash. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Oliver Nash +-/ +import linear_algebra.affine_space.basis +import linear_algebra.determinant + +/-! +# Matrix results for barycentric co-ordinates + +Results about the matrix of barycentric co-ordinates for a family of points in an affine space, with +respect to some affine basis. +-/ + +open_locale affine big_operators matrix +open set + +universes u₁ u₂ u₃ u₄ + +variables {ι : Type u₁} {k : Type u₂} {V : Type u₃} {P : Type u₄} +variables [add_comm_group V] [affine_space V P] + +namespace affine_basis + +section ring + +variables [ring k] [module k V] (b : affine_basis ι k P) + +/-- Given an affine basis `p`, and a family of points `q : ι' → P`, this is the matrix whose +rows are the barycentric coordinates of `q` with respect to `p`. + +It is an affine equivalent of `basis.to_matrix`. -/ +noncomputable def to_matrix {ι' : Type*} (q : ι' → P) : matrix ι' ι k := +λ i j, b.coord j (q i) + +@[simp] lemma to_matrix_apply {ι' : Type*} (q : ι' → P) (i : ι') (j : ι) : + b.to_matrix q i j = b.coord j (q i) := +rfl + +@[simp] lemma to_matrix_self [decidable_eq ι] : + b.to_matrix b.points = (1 : matrix ι ι k) := +begin + ext i j, + rw [to_matrix_apply, coord_apply, matrix.one_eq_pi_single, pi.single_apply], +end + +variables {ι' : Type*} [fintype ι'] [fintype ι] (b₂ : affine_basis ι k P) + +lemma to_matrix_row_sum_one {ι' : Type*} (q : ι' → P) (i : ι') : + ∑ j, b.to_matrix q i j = 1 := +by simp + +/-- Given a family of points `p : ι' → P` and an affine basis `b`, if the matrix whose rows are the +coordinates of `p` with respect `b` has a right inverse, then `p` is affine independent. -/ +lemma affine_independent_of_to_matrix_right_inv [decidable_eq ι'] + (p : ι' → P) {A : matrix ι ι' k} (hA : (b.to_matrix p) ⬝ A = 1) : affine_independent k p := +begin + rw affine_independent_iff_eq_of_fintype_affine_combination_eq, + intros w₁ w₂ hw₁ hw₂ hweq, + have hweq' : (b.to_matrix p).vec_mul w₁ = (b.to_matrix p).vec_mul w₂, + { ext j, + change ∑ i, (w₁ i) • (b.coord j (p i)) = ∑ i, (w₂ i) • (b.coord j (p i)), + rw [← finset.univ.affine_combination_eq_linear_combination _ _ hw₁, + ← finset.univ.affine_combination_eq_linear_combination _ _ hw₂, + ← finset.univ.map_affine_combination p w₁ hw₁, + ← finset.univ.map_affine_combination p w₂ hw₂, hweq], }, + replace hweq' := congr_arg (λ w, A.vec_mul w) hweq', + simpa only [matrix.vec_mul_vec_mul, ← matrix.mul_eq_mul, hA, matrix.vec_mul_one] using hweq', +end + +/-- Given a family of points `p : ι' → P` and an affine basis `b`, if the matrix whose rows are the +coordinates of `p` with respect `b` has a left inverse, then `p` spans the entire space. -/ +lemma affine_span_eq_top_of_to_matrix_left_inv [decidable_eq ι] [nontrivial k] + (p : ι' → P) {A : matrix ι ι' k} (hA : A ⬝ b.to_matrix p = 1) : affine_span k (range p) = ⊤ := +begin + suffices : ∀ i, b.points i ∈ affine_span k (range p), + { rw [eq_top_iff, ← b.tot, affine_span_le], + rintros q ⟨i, rfl⟩, + exact this i, }, + intros i, + have hAi : ∑ j, A i j = 1, + { calc ∑ j, A i j = ∑ j, (A i j) * ∑ l, b.to_matrix p j l : by simp + ... = ∑ j, ∑ l, (A i j) * b.to_matrix p j l : by simp_rw finset.mul_sum + ... = ∑ l, ∑ j, (A i j) * b.to_matrix p j l : by rw finset.sum_comm + ... = ∑ l, (A ⬝ b.to_matrix p) i l : rfl + ... = 1 : by simp [hA, matrix.one_apply, finset.filter_eq], }, + have hbi : b.points i = finset.univ.affine_combination p (A i), + { apply b.ext_elem, + intros j, + rw [b.coord_apply, finset.univ.map_affine_combination _ _ hAi, + finset.univ.affine_combination_eq_linear_combination _ _ hAi], + change _ = (A ⬝ b.to_matrix p) i j, + simp_rw [hA, matrix.one_apply, @eq_comm _ i j] }, + rw hbi, + exact affine_combination_mem_affine_span hAi p, +end + +/-- A change of basis formula for barycentric coordinates. + +See also `affine_basis.to_matrix_inv_mul_affine_basis_to_matrix`. -/ +@[simp] lemma to_matrix_vec_mul_coords (x : P) : + (b.to_matrix b₂.points).vec_mul (b₂.coords x) = b.coords x := +begin + ext j, + change _ = b.coord j x, + conv_rhs { rw ← b₂.affine_combination_coord_eq_self x, }, + rw finset.map_affine_combination _ _ _ (b₂.sum_coord_apply_eq_one x), + simp [matrix.vec_mul, matrix.dot_product, to_matrix_apply, coords], +end + +variables [decidable_eq ι] + +lemma to_matrix_mul_to_matrix : + (b.to_matrix b₂.points) ⬝ (b₂.to_matrix b.points) = 1 := +begin + ext l m, + change (b₂.to_matrix b.points).vec_mul (b.coords (b₂.points l)) m = _, + rw [to_matrix_vec_mul_coords, coords_apply, ← to_matrix_apply, to_matrix_self], +end + +lemma is_unit_to_matrix : + is_unit (b.to_matrix b₂.points) := +⟨{ val := b.to_matrix b₂.points, + inv := b₂.to_matrix b.points, + val_inv := b.to_matrix_mul_to_matrix b₂, + inv_val := b₂.to_matrix_mul_to_matrix b, }, rfl⟩ + +lemma is_unit_to_matrix_iff [nontrivial k] (p : ι → P) : + is_unit (b.to_matrix p) ↔ affine_independent k p ∧ affine_span k (range p) = ⊤ := +begin + split, + { rintros ⟨⟨B, A, hA, hA'⟩, (rfl : B = b.to_matrix p)⟩, + rw matrix.mul_eq_mul at hA hA', + exact ⟨b.affine_independent_of_to_matrix_right_inv p hA, + b.affine_span_eq_top_of_to_matrix_left_inv p hA'⟩, }, + { rintros ⟨h_tot, h_ind⟩, + let b' : affine_basis ι k P := ⟨p, h_tot, h_ind⟩, + change is_unit (b.to_matrix b'.points), + exact b.is_unit_to_matrix b', }, +end + +end ring + +section comm_ring +variables [comm_ring k] [module k V] [decidable_eq ι] [fintype ι] +variables (b b₂ : affine_basis ι k P) + +/-- A change of basis formula for barycentric coordinates. + +See also `affine_basis.to_matrix_vec_mul_coords`. -/ +@[simp] lemma to_matrix_inv_vec_mul_to_matrix (x : P) : + (b.to_matrix b₂.points)⁻¹.vec_mul (b.coords x) = b₂.coords x := +begin + have hu := b.is_unit_to_matrix b₂, + rw matrix.is_unit_iff_is_unit_det at hu, + rw [← b.to_matrix_vec_mul_coords b₂, matrix.vec_mul_vec_mul, matrix.mul_nonsing_inv _ hu, + matrix.vec_mul_one], +end + +/-- If we fix a background affine basis `b`, then for any other basis `b₂`, we can characterise +the barycentric coordinates provided by `b₂` in terms of determinants relative to `b`. -/ +lemma det_smul_coords_eq_cramer_coords (x : P) : + (b.to_matrix b₂.points).det • b₂.coords x = (b.to_matrix b₂.points)ᵀ.cramer (b.coords x) := +begin + have hu := b.is_unit_to_matrix b₂, + rw matrix.is_unit_iff_is_unit_det at hu, + rw [← b.to_matrix_inv_vec_mul_to_matrix, matrix.det_smul_inv_vec_mul_eq_cramer_transpose _ _ hu], +end +end comm_ring + +end affine_basis From 008af8bb14b3ebef7e04ec3b0d63b947dee4d26a Mon Sep 17 00:00:00 2001 From: negiizhao Date: Thu, 12 Jan 2023 11:00:11 +0000 Subject: [PATCH 0219/1291] chore(data/fin/basic): remove `fin.coe_of_nat_eq_mod` (#18131) It can be proved by `fin.coe_of_nat_eq_mod'`. --- src/data/bitvec/basic.lean | 4 ++-- src/data/fin/basic.lean | 7 +------ 2 files changed, 3 insertions(+), 8 deletions(-) diff --git a/src/data/bitvec/basic.lean b/src/data/bitvec/basic.lean index c5715812f85be..5e76ec218b003 100644 --- a/src/data/bitvec/basic.lean +++ b/src/data/bitvec/basic.lean @@ -22,7 +22,7 @@ by rw [of_fin,to_nat_of_nat,nat.mod_eq_of_lt]; apply i.is_lt /-- convert `bitvec` to `fin` -/ def to_fin {n : ℕ} (i : bitvec n) : fin $ 2^n := -fin.of_nat' i.to_nat +i.to_nat lemma add_lsb_eq_twice_add_one {x b} : add_lsb x b = 2 * x + cond b 1 0 := @@ -77,7 +77,7 @@ begin end lemma to_fin_val {n : ℕ} (v : bitvec n) : (to_fin v : ℕ) = v.to_nat := -by rw [to_fin, fin.of_nat'_eq_coe, fin.coe_of_nat_eq_mod', nat.mod_eq_of_lt]; apply to_nat_lt +by rw [to_fin, fin.coe_of_nat_eq_mod, nat.mod_eq_of_lt]; apply to_nat_lt lemma to_fin_le_to_fin_of_le {n} {v₀ v₁ : bitvec n} (h : v₀ ≤ v₁) : v₀.to_fin ≤ v₁.to_fin := show (v₀.to_fin : ℕ) ≤ v₁.to_fin, diff --git a/src/data/fin/basic.lean b/src/data/fin/basic.lean index aae2f9ffddc75..2be427042b554 100644 --- a/src/data/fin/basic.lean +++ b/src/data/fin/basic.lean @@ -1866,15 +1866,10 @@ def clamp (n m : ℕ) : fin (m + 1) := of_nat $ min n m @[simp] lemma coe_clamp (n m : ℕ) : (clamp n m : ℕ) = min n m := nat.mod_eq_of_lt $ nat.lt_succ_iff.mpr $ min_le_right _ _ -@[simp] lemma coe_of_nat_eq_mod' (m n : ℕ) [I : ne_zero m] : +@[simp] lemma coe_of_nat_eq_mod (m n : ℕ) [ne_zero m] : ((n : fin m) : ℕ) = n % m := rfl -@[simp] -lemma coe_of_nat_eq_mod (m n : ℕ) : - ((n : fin (succ m)) : ℕ) = n % succ m := -by rw [← of_nat_eq_coe]; refl - section mul /-! From b0c712376e4ef44c53c3b872157ef44dfe9f9599 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 12 Jan 2023 11:00:13 +0000 Subject: [PATCH 0220/1291] chore(linear_algebra/tensor_product): make more lemmas dsimp-eligible (#18140) These are more items to add to the list in #18121 --- src/linear_algebra/dual.lean | 9 +++------ src/linear_algebra/tensor_product.lean | 3 +-- 2 files changed, 4 insertions(+), 8 deletions(-) diff --git a/src/linear_algebra/dual.lean b/src/linear_algebra/dual.lean index 98e9fc9892c8e..425bd69a89e60 100644 --- a/src/linear_algebra/dual.lean +++ b/src/linear_algebra/dual.lean @@ -780,8 +780,7 @@ def linear_map.dual_map (f : M₁ →ₗ[R] M₂) : dual R M₂ →ₗ[R] dual R linear_map.lcomp R R f @[simp] lemma linear_map.dual_map_apply (f : M₁ →ₗ[R] M₂) (g : dual R M₂) (x : M₁) : - f.dual_map g x = g (f x) := -linear_map.lcomp_apply f g x + f.dual_map g x = g (f x) := rfl @[simp] lemma linear_map.dual_map_id : (linear_map.id : M₁ →ₗ[R] M₁).dual_map = linear_map.id := @@ -810,8 +809,7 @@ def linear_equiv.dual_map (f : M₁ ≃ₗ[R] M₂) : dual R M₂ ≃ₗ[R] dual .. f.to_linear_map.dual_map } @[simp] lemma linear_equiv.dual_map_apply (f : M₁ ≃ₗ[R] M₂) (g : dual R M₂) (x : M₁) : - f.dual_map g x = g (f x) := -linear_map.lcomp_apply f g x + f.dual_map g x = g (f x) := rfl @[simp] lemma linear_equiv.dual_map_refl : (linear_equiv.refl R M₁).dual_map = linear_equiv.refl R (dual R M₁) := @@ -947,8 +945,7 @@ variables {R M N} @[simp] lemma dual_distrib_apply (f : dual R M) (g : dual R N) (m : M) (n : N) : dual_distrib R M N (f ⊗ₜ g) (m ⊗ₜ n) = f m * g n := -by simp only [dual_distrib, coe_comp, function.comp_app, hom_tensor_hom_map_apply, - comp_right_apply, linear_equiv.coe_coe, map_tmul, lid_tmul, algebra.id.smul_eq_mul] +rfl end diff --git a/src/linear_algebra/tensor_product.lean b/src/linear_algebra/tensor_product.lean index f0a423cee2d83..eb1724d4b5006 100644 --- a/src/linear_algebra/tensor_product.lean +++ b/src/linear_algebra/tensor_product.lean @@ -721,8 +721,7 @@ lemma rtensor_hom_to_hom_rtensor_apply (f : M →ₗ[R] P) (q : Q) (m : M) : @[simp] lemma hom_tensor_hom_map_apply (f : M →ₗ[R] P) (g : N →ₗ[R] Q) : - hom_tensor_hom_map R M N P Q (f ⊗ₜ g) = map f g := -by simp only [hom_tensor_hom_map, lift.tmul, map_bilinear_apply] + hom_tensor_hom_map R M N P Q (f ⊗ₜ g) = map f g := rfl end From 217daaf45adff6a74a82df0dab1506ed3d5e2eec Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Thu, 12 Jan 2023 14:20:25 +0000 Subject: [PATCH 0221/1291] feat(geometry/euclidean/angle/unoriented/affine): more collinearity lemmas (#18069) These lemmas relating collinearity and angles of zero or pi are less general than the existing `collinear_iff_eq_or_eq_or_angle_eq_zero_or_angle_eq_pi` from which they are deduced, but may be more convenient to use. Suggested by review of #17993. --- .../euclidean/angle/unoriented/affine.lean | 30 +++++++++++++++++++ 1 file changed, 30 insertions(+) diff --git a/src/geometry/euclidean/angle/unoriented/affine.lean b/src/geometry/euclidean/angle/unoriented/affine.lean index e99e34421fa34..c273b56f61253 100644 --- a/src/geometry/euclidean/angle/unoriented/affine.lean +++ b/src/geometry/euclidean/angle/unoriented/affine.lean @@ -409,4 +409,34 @@ begin exact h.wbtw.collinear } } end +/-- If the angle between three points is 0, they are collinear. -/ +lemma collinear_of_angle_eq_zero {p₁ p₂ p₃ : P} (h : ∠ p₁ p₂ p₃ = 0) : + collinear ℝ ({p₁, p₂, p₃} : set P) := +collinear_iff_eq_or_eq_or_angle_eq_zero_or_angle_eq_pi.2 $ or.inr $ or.inr $ or.inl h + +/-- If the angle between three points is π, they are collinear. -/ +lemma collinear_of_angle_eq_pi {p₁ p₂ p₃ : P} (h : ∠ p₁ p₂ p₃ = π) : + collinear ℝ ({p₁, p₂, p₃} : set P) := +collinear_iff_eq_or_eq_or_angle_eq_zero_or_angle_eq_pi.2 $ or.inr $ or.inr $ or.inr h + +/-- If three points are not collinear, the angle between them is nonzero. -/ +lemma angle_ne_zero_of_not_collinear {p₁ p₂ p₃ : P} (h : ¬collinear ℝ ({p₁, p₂, p₃} : set P)) : + ∠ p₁ p₂ p₃ ≠ 0 := +mt collinear_of_angle_eq_zero h + +/-- If three points are not collinear, the angle between them is not π. -/ +lemma angle_ne_pi_of_not_collinear {p₁ p₂ p₃ : P} (h : ¬collinear ℝ ({p₁, p₂, p₃} : set P)) : + ∠ p₁ p₂ p₃ ≠ π := +mt collinear_of_angle_eq_pi h + +/-- If three points are not collinear, the angle between them is positive. -/ +lemma angle_pos_of_not_collinear {p₁ p₂ p₃ : P} (h : ¬collinear ℝ ({p₁, p₂, p₃} : set P)) : + 0 < ∠ p₁ p₂ p₃ := +(angle_nonneg _ _ _).lt_of_ne (angle_ne_zero_of_not_collinear h).symm + +/-- If three points are not collinear, the angle between them is less than π. -/ +lemma angle_lt_pi_of_not_collinear {p₁ p₂ p₃ : P} (h : ¬collinear ℝ ({p₁, p₂, p₃} : set P)) : + ∠ p₁ p₂ p₃ < π := +(angle_le_pi _ _ _).lt_of_ne $ angle_ne_pi_of_not_collinear h + end euclidean_geometry From 1db04677dde71458ada63faa8336054f24490ad7 Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Thu, 12 Jan 2023 15:42:10 +0000 Subject: [PATCH 0222/1291] feat(data/polynomial): `adjoin_root.mk_ne_zero` lemmas (#18113) + `mk_ne_zero_of_(nat_)degree_lt`: if `f : R[X]` is monic and `g : R[X]` is nonzero with (nat_)degree less than (nat_)degree of `f`, then the image of `g` in `R[X]/(f)` is nonzero. Depends on new lemmas `monic.not_dvd_of_(nat_)degree_lt`. + Use it to golf `weierstrass_curve.X/Y_class_ne_zero`. + Move misplaced lemma `nat_degree_lt_nat_degree`. --- .../elliptic_curve/weierstrass.lean | 18 +++---------- src/data/polynomial/degree/definitions.lean | 7 +++++ src/data/polynomial/monic.lean | 26 ++++++++++++++----- src/ring_theory/adjoin_root.lean | 11 ++++++++ src/ring_theory/power_basis.lean | 11 +------- 5 files changed, 42 insertions(+), 31 deletions(-) diff --git a/src/algebraic_geometry/elliptic_curve/weierstrass.lean b/src/algebraic_geometry/elliptic_curve/weierstrass.lean index 5e1e64ed890f9..15703ad46c5da 100644 --- a/src/algebraic_geometry/elliptic_curve/weierstrass.lean +++ b/src/algebraic_geometry/elliptic_curve/weierstrass.lean @@ -400,25 +400,15 @@ variables (x : R) (y : R[X]) @[simp] noncomputable def X_class : W.coordinate_ring := adjoin_root.mk W.polynomial $ C $ X - C x lemma X_class_ne_zero [nontrivial R] : W.X_class x ≠ 0 := -begin - intro hx, - cases ideal.mem_span_singleton'.mp (ideal.quotient.eq_zero_iff_mem.mp hx) with p hp, - apply_fun degree at hp, - rw [W.monic_polynomial.degree_mul, degree_polynomial, degree_C $ X_sub_C_ne_zero x] at hp, - cases p.degree; cases hp -end +adjoin_root.mk_ne_zero_of_nat_degree_lt W.monic_polynomial (C_ne_zero.2 $ X_sub_C_ne_zero x) $ + by { rw [nat_degree_polynomial, nat_degree_C], norm_num1 } /-- The class of the element $Y - y(X)$ in $R[W]$ for some $y(X) \in R[X]$. -/ @[simp] noncomputable def Y_class : W.coordinate_ring := adjoin_root.mk W.polynomial $ X - C y lemma Y_class_ne_zero [nontrivial R] : W.Y_class y ≠ 0 := -begin - intro hy, - cases ideal.mem_span_singleton'.mp (ideal.quotient.eq_zero_iff_mem.mp hy) with p hp, - apply_fun degree at hp, - rw [W.monic_polynomial.degree_mul, degree_polynomial, degree_X_sub_C] at hp, - cases p.degree; cases hp -end +adjoin_root.mk_ne_zero_of_nat_degree_lt W.monic_polynomial (X_sub_C_ne_zero _) $ + by { rw [nat_degree_polynomial, nat_degree_X_sub_C], norm_num1 } /-- The ideal $\langle X - x \rangle$ of $R[W]$ for some $x \in R$. -/ @[simp] noncomputable def X_ideal : ideal W.coordinate_ring := ideal.span {W.X_class x} diff --git a/src/data/polynomial/degree/definitions.lean b/src/data/polynomial/degree/definitions.lean index 41c44adb355f5..bac87138058bb 100644 --- a/src/data/polynomial/degree/definitions.lean +++ b/src/data/polynomial/degree/definitions.lean @@ -169,6 +169,13 @@ lemma nat_degree_le_nat_degree [semiring S] {q : S[X]} (hpq : p.degree ≤ q.deg p.nat_degree ≤ q.nat_degree := with_bot.gi_unbot'_bot.gc.monotone_l hpq +lemma nat_degree_lt_nat_degree {p q : R[X]} (hp : p ≠ 0) (hpq : p.degree < q.degree) : + p.nat_degree < q.nat_degree := +begin + by_cases hq : q = 0, { exact (not_lt_bot $ hq.subst hpq).elim }, + rwa [degree_eq_nat_degree hp, degree_eq_nat_degree hq, with_bot.coe_lt_coe] at hpq +end + @[simp] lemma degree_C (ha : a ≠ 0) : degree (C a) = (0 : with_bot ℕ) := by rw [degree, ← monomial_zero_left, support_monomial 0 ha, max_eq_sup_coe, sup_singleton, with_bot.coe_zero] diff --git a/src/data/polynomial/monic.lean b/src/data/polynomial/monic.lean index da99d33616e88..77cad7c1143cf 100644 --- a/src/data/polynomial/monic.lean +++ b/src/data/polynomial/monic.lean @@ -125,7 +125,7 @@ end namespace monic @[simp] -lemma nat_degree_eq_zero_iff_eq_one {p : R[X]} (hp : p.monic) : +lemma nat_degree_eq_zero_iff_eq_one (hp : p.monic) : p.nat_degree = 0 ↔ p = 1 := begin split; intro h, @@ -137,11 +137,11 @@ begin end @[simp] -lemma degree_le_zero_iff_eq_one {p : R[X]} (hp : p.monic) : +lemma degree_le_zero_iff_eq_one (hp : p.monic) : p.degree ≤ 0 ↔ p = 1 := by rw [←hp.nat_degree_eq_zero_iff_eq_one, nat_degree_eq_zero_iff_degree_le_zero] -lemma nat_degree_mul {p q : R[X]} (hp : p.monic) (hq : q.monic) : +lemma nat_degree_mul (hp : p.monic) (hq : q.monic) : (p * q).nat_degree = p.nat_degree + q.nat_degree := begin nontriviality R, @@ -149,7 +149,7 @@ begin simp [hp.leading_coeff, hq.leading_coeff] end -lemma degree_mul_comm {p : R[X]} (hp : p.monic) (q : R[X]) : +lemma degree_mul_comm (hp : p.monic) (q : R[X]) : (p * q).degree = (q * p).degree := begin by_cases h : q = 0, @@ -159,14 +159,14 @@ begin { rwa [hp.leading_coeff, one_mul, leading_coeff_ne_zero] } end -lemma nat_degree_mul' {p q : R[X]} (hp : p.monic) (hq : q ≠ 0) : +lemma nat_degree_mul' (hp : p.monic) (hq : q ≠ 0) : (p * q).nat_degree = p.nat_degree + q.nat_degree := begin rw [nat_degree_mul', add_comm], simpa [hp.leading_coeff, leading_coeff_ne_zero] end -lemma nat_degree_mul_comm {p : R[X]} (hp : p.monic) (q : R[X]) : +lemma nat_degree_mul_comm (hp : p.monic) (q : R[X]) : (p * q).nat_degree = (q * p).nat_degree := begin by_cases h : q = 0, @@ -175,7 +175,19 @@ begin simpa [hp.leading_coeff, leading_coeff_ne_zero] end -lemma next_coeff_mul {p q : R[X]} (hp : monic p) (hq : monic q) : +lemma not_dvd_of_nat_degree_lt (hp : monic p) + (h0 : q ≠ 0) (hl : nat_degree q < nat_degree p) : ¬ p ∣ q := +begin + rintro ⟨r, rfl⟩, + rw [hp.nat_degree_mul' $ right_ne_zero_of_mul h0] at hl, + exact hl.not_le (nat.le_add_right _ _) +end + +lemma not_dvd_of_degree_lt (hp : monic p) + (h0 : q ≠ 0) (hl : degree q < degree p) : ¬ p ∣ q := +monic.not_dvd_of_nat_degree_lt hp h0 $ nat_degree_lt_nat_degree h0 hl + +lemma next_coeff_mul (hp : monic p) (hq : monic q) : next_coeff (p * q) = next_coeff p + next_coeff q := begin nontriviality, diff --git a/src/ring_theory/adjoin_root.lean b/src/ring_theory/adjoin_root.lean index a99085c2611d8..d3f088f947c3c 100644 --- a/src/ring_theory/adjoin_root.lean +++ b/src/ring_theory/adjoin_root.lean @@ -131,6 +131,9 @@ ideal.quotient.alg_hom_ext R $ polynomial.alg_hom_ext h @[simp] lemma mk_eq_mk {g h : R[X]} : mk f g = mk f h ↔ f ∣ g - h := ideal.quotient.eq.trans ideal.mem_span_singleton +@[simp] lemma mk_eq_zero {g : R[X]} : mk f g = 0 ↔ f ∣ g := +mk_eq_mk.trans $ by rw sub_zero + @[simp] lemma mk_self : mk f f = 0 := quotient.sound' $ quotient_add_group.left_rel_apply.mpr (mem_span_singleton.2 $ by simp) @@ -138,6 +141,14 @@ quotient.sound' $ quotient_add_group.left_rel_apply.mpr (mem_span_singleton.2 $ @[simp] lemma mk_X : mk f X = root f := rfl +lemma mk_ne_zero_of_degree_lt (hf : monic f) + {g : R[X]} (h0 : g ≠ 0) (hd : degree g < degree f) : mk f g ≠ 0 := +mk_eq_zero.not.2 $ hf.not_dvd_of_degree_lt h0 hd + +lemma mk_ne_zero_of_nat_degree_lt (hf : monic f) + {g : R[X]} (h0 : g ≠ 0) (hd : nat_degree g < nat_degree f) : mk f g ≠ 0 := +mk_eq_zero.not.2 $ hf.not_dvd_of_nat_degree_lt h0 hd + @[simp] lemma aeval_eq (p : R[X]) : aeval (root f) p = mk f p := polynomial.induction_on p (λ x, by { rw aeval_C, refl }) (λ p q ihp ihq, by rw [alg_hom.map_add, ring_hom.map_add, ihp, ihq]) diff --git a/src/ring_theory/power_basis.lean b/src/ring_theory/power_basis.lean index fb053509a8d99..a1f0fdba7af3c 100644 --- a/src/ring_theory/power_basis.lean +++ b/src/ring_theory/power_basis.lean @@ -239,16 +239,7 @@ end minpoly section equiv -variables [algebra A S] {S' : Type*} [comm_ring S'] [algebra A S'] - -lemma nat_degree_lt_nat_degree {p q : R[X]} (hp : p ≠ 0) (hpq : p.degree < q.degree) : - p.nat_degree < q.nat_degree := -begin - by_cases hq : q = 0, { rw [hq, degree_zero] at hpq, have := not_lt_bot hpq, contradiction }, - rwa [degree_eq_nat_degree hp, degree_eq_nat_degree hq, with_bot.coe_lt_coe] at hpq -end - -variables [is_domain A] +variables [algebra A S] {S' : Type*} [comm_ring S'] [algebra A S'] [is_domain A] lemma constr_pow_aeval (pb : power_basis A S) {y : S'} (hy : aeval y (minpoly A pb.gen) = 0) (f : A[X]) : From 9003f28797c0664a49e4179487267c494477d853 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 12 Jan 2023 17:56:13 +0000 Subject: [PATCH 0223/1291] chore(data/set/intervals/unordered_interval): Rename to `uIcc`/`uIoc` (#18104) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Rename * `set.interval` → `set.uIcc` * `set.interval_oc` → `set.uIoc` * `finset.interval`→ `finset.uIcc` Closes #17982 Zulip: https://leanprover.zulipchat.com/#narrow/stream/116395-maths/topic/Set.20intervals.20names --- src/analysis/ODE/picard_lindelof.lean | 6 +- .../parametric_interval_integral.lean | 16 +- src/analysis/complex/cauchy_integral.lean | 4 +- src/analysis/convex/basic.lean | 5 +- src/analysis/convex/quasiconvex.lean | 6 +- src/analysis/convex/segment.lean | 12 +- src/analysis/convex/star.lean | 3 +- src/analysis/convex/strict.lean | 6 +- src/analysis/special_functions/integrals.lean | 26 +-- .../special_functions/non_integrable.lean | 8 +- src/analysis/sum_integral_comparisons.lean | 4 +- src/data/finset/locally_finite.lean | 93 ++++---- src/data/set/intervals/ord_connected.lean | 30 +-- .../intervals/ord_connected_component.lean | 37 ++-- .../set/intervals/unordered_interval.lean | 199 ++++++++---------- src/data/set/pointwise/interval.lean | 56 ++--- .../affine_space/affine_map.lean | 6 +- .../constructions/borel_space.lean | 13 +- .../function/locally_integrable.lean | 6 +- .../function/strongly_measurable/basic.lean | 6 +- .../integral/circle_integral.lean | 2 +- .../integral/circle_transform.lean | 6 +- .../integral/divergence_theorem.lean | 10 +- .../integral/integral_eq_improper.lean | 4 +- .../integral/interval_average.lean | 6 +- .../integral/interval_integral.lean | 196 ++++++++--------- src/measure_theory/integral/layercake.lean | 2 +- src/measure_theory/measure/ae_measurable.lean | 4 +- src/measure_theory/measure/lebesgue.lean | 2 +- src/measure_theory/measure/measure_space.lean | 18 +- src/order/filter/interval.lean | 14 +- src/order/locally_finite.lean | 28 +-- src/topology/algebra/order/compact.lean | 16 +- .../algebra/order/intermediate_value.lean | 16 +- src/topology/algebra/order/t5.lean | 12 +- src/topology/metric_space/basic.lean | 20 +- 36 files changed, 432 insertions(+), 466 deletions(-) diff --git a/src/analysis/ODE/picard_lindelof.lean b/src/analysis/ODE/picard_lindelof.lean index 4e670e38341cb..a7bfe958abd54 100644 --- a/src/analysis/ODE/picard_lindelof.lean +++ b/src/analysis/ODE/picard_lindelof.lean @@ -265,17 +265,17 @@ begin calc ‖∫ τ in Ι (v.t₀ : ℝ) t, f₁.v_comp τ - f₂.v_comp τ‖ ≤ ∫ τ in Ι (v.t₀ : ℝ) t, v.L * ((v.L * |τ - v.t₀|) ^ n / n! * d) : begin - refine norm_integral_le_of_norm_le (continuous.integrable_on_interval_oc _) _, + refine norm_integral_le_of_norm_le (continuous.integrable_on_uIoc _) _, { continuity }, { refine (ae_restrict_mem measurable_set_Ioc).mono (λ τ hτ, _), refine (v.lipschitz_on_with (v.proj τ).2).norm_sub_le_of_le (f₁.mem_closed_ball _) (f₂.mem_closed_ball _) ((h _).trans_eq _), rw v.proj_of_mem, - exact (interval_subset_Icc v.t₀.2 t.2 $ Ioc_subset_Icc_self hτ) } + exact (uIcc_subset_Icc v.t₀.2 t.2 $ Ioc_subset_Icc_self hτ) } end ... = (v.L * |t - v.t₀|) ^ (n + 1) / (n + 1)! * d : _, simp_rw [mul_pow, div_eq_mul_inv, mul_assoc, measure_theory.integral_mul_left, - measure_theory.integral_mul_right, integral_pow_abs_sub_interval_oc, div_eq_mul_inv, + measure_theory.integral_mul_right, integral_pow_abs_sub_uIoc, div_eq_mul_inv, pow_succ (v.L : ℝ), nat.factorial_succ, nat.cast_mul, nat.cast_succ, mul_inv, mul_assoc] end diff --git a/src/analysis/calculus/parametric_interval_integral.lean b/src/analysis/calculus/parametric_interval_integral.lean index 03a19aa2a3649..f6e82e4cef673 100644 --- a/src/analysis/calculus/parametric_interval_integral.lean +++ b/src/analysis/calculus/parametric_interval_integral.lean @@ -39,8 +39,8 @@ lemma has_fderiv_at_integral_of_dominated_loc_of_lip {F : H → ℝ → E} {F' : interval_integrable F' μ a b ∧ has_fderiv_at (λ x, ∫ t in a..b, F x t ∂μ) (∫ t in a..b, F' t ∂μ) x₀ := begin - simp only [interval_integrable_iff, interval_integral_eq_integral_interval_oc, - ← ae_restrict_iff' measurable_set_interval_oc] at *, + simp only [interval_integrable_iff, interval_integral_eq_integral_uIoc, + ← ae_restrict_iff' measurable_set_uIoc] at *, have := has_fderiv_at_integral_of_dominated_loc_of_lip ε_pos hF_meas hF_int hF'_meas h_lip bound_integrable h_diff, exact ⟨this.1, this.2.const_smul _⟩ @@ -60,8 +60,8 @@ lemma has_fderiv_at_integral_of_dominated_of_fderiv_le {F : H → ℝ → E} {F' (h_diff : ∀ᵐ t ∂μ, t ∈ Ι a b → ∀ x ∈ ball x₀ ε, has_fderiv_at (λ x, F x t) (F' x t) x) : has_fderiv_at (λ x, ∫ t in a..b, F x t ∂μ) (∫ t in a..b, F' x₀ t ∂μ) x₀ := begin - simp only [interval_integrable_iff, interval_integral_eq_integral_interval_oc, - ← ae_restrict_iff' measurable_set_interval_oc] at *, + simp only [interval_integrable_iff, interval_integral_eq_integral_uIoc, + ← ae_restrict_iff' measurable_set_uIoc] at *, exact (has_fderiv_at_integral_of_dominated_of_fderiv_le ε_pos hF_meas hF_int hF'_meas h_bound bound_integrable h_diff).const_smul _ end @@ -82,8 +82,8 @@ lemma has_deriv_at_integral_of_dominated_loc_of_lip {F : 𝕜 → ℝ → E} {F' (interval_integrable F' μ a b) ∧ has_deriv_at (λ x, ∫ t in a..b, F x t ∂μ) (∫ t in a..b, F' t ∂μ) x₀ := begin - simp only [interval_integrable_iff, interval_integral_eq_integral_interval_oc, - ← ae_restrict_iff' measurable_set_interval_oc] at *, + simp only [interval_integrable_iff, interval_integral_eq_integral_uIoc, + ← ae_restrict_iff' measurable_set_uIoc] at *, have := has_deriv_at_integral_of_dominated_loc_of_lip ε_pos hF_meas hF_int hF'_meas h_lipsch bound_integrable h_diff, exact ⟨this.1, this.2.const_smul _⟩ @@ -104,8 +104,8 @@ lemma has_deriv_at_integral_of_dominated_loc_of_deriv_le {F : 𝕜 → ℝ → E (interval_integrable (F' x₀) μ a b) ∧ has_deriv_at (λ x, ∫ t in a..b, F x t ∂μ) (∫ t in a..b, F' x₀ t ∂μ) x₀ := begin - simp only [interval_integrable_iff, interval_integral_eq_integral_interval_oc, - ← ae_restrict_iff' measurable_set_interval_oc] at *, + simp only [interval_integrable_iff, interval_integral_eq_integral_uIoc, + ← ae_restrict_iff' measurable_set_uIoc] at *, have := has_deriv_at_integral_of_dominated_loc_of_deriv_le ε_pos hF_meas hF_int hF'_meas h_bound bound_integrable h_diff, exact ⟨this.1, this.2.const_smul _⟩ diff --git a/src/analysis/complex/cauchy_integral.lean b/src/analysis/complex/cauchy_integral.lean index d1605f7aad8f7..2cbc77d447e99 100644 --- a/src/analysis/complex/cauchy_integral.lean +++ b/src/analysis/complex/cauchy_integral.lean @@ -179,7 +179,7 @@ begin { rintro ⟨x, y⟩, simp [F', he₁, he₂, ← sub_eq_neg_add], }, set R : set (ℝ × ℝ) := [z.re, w.re] ×ˢ [w.im, z.im], set t : set (ℝ × ℝ) := e ⁻¹' s, - rw [interval_swap z.im] at Hc Hi, rw [min_comm z.im, max_comm z.im] at Hd, + rw [uIcc_comm z.im] at Hc Hi, rw [min_comm z.im, max_comm z.im] at Hd, have hR : e ⁻¹' ([z.re, w.re] ×ℂ [w.im, z.im]) = R := rfl, have htc : continuous_on F R, from Hc.comp e.continuous_on hR.ge, have htd : ∀ p ∈ Ioo (min z.re w.re) (max z.re w.re) ×ˢ Ioo (min w.im z.im) (max w.im z.im) \ t, @@ -227,7 +227,7 @@ lemma integral_boundary_rect_of_differentiable_on_real (f : ℂ → E) (z w : integral_boundary_rect_of_has_fderiv_at_real_off_countable f (fderiv ℝ f) z w ∅ countable_empty Hd.continuous_on (λ x hx, Hd.has_fderiv_at $ by simpa only [← mem_interior_iff_mem_nhds, - interior_re_prod_im, interval, interior_Icc] using hx.1) Hi + interior_re_prod_im, uIcc, interior_Icc] using hx.1) Hi /-- **Cauchy-Goursat theorem** for a rectangle: the integral of a complex differentiable function over the boundary of a rectangle equals zero. More precisely, if `f` is continuous on a closed diff --git a/src/analysis/convex/basic.lean b/src/analysis/convex/basic.lean index a3d0acf4558cf..601c2623ecfe3 100644 --- a/src/analysis/convex/basic.lean +++ b/src/analysis/convex/basic.lean @@ -278,8 +278,7 @@ end ordered_cancel_add_comm_monoid section linear_ordered_add_comm_monoid variables [linear_ordered_add_comm_monoid β] [module 𝕜 β] [ordered_smul 𝕜 β] -lemma convex_interval (r s : β) : convex 𝕜 (interval r s) := -convex_Icc _ _ +lemma convex_uIcc (r s : β) : convex 𝕜 (uIcc r s) := convex_Icc _ _ end linear_ordered_add_comm_monoid end module @@ -513,7 +512,7 @@ hs.convex_of_chain $ is_chain_of_trichotomous s lemma convex_iff_ord_connected [linear_ordered_field 𝕜] {s : set 𝕜} : convex 𝕜 s ↔ s.ord_connected := -by simp_rw [convex_iff_segment_subset, segment_eq_interval, ord_connected_iff_interval_subset] +by simp_rw [convex_iff_segment_subset, segment_eq_uIcc, ord_connected_iff_uIcc_subset] alias convex_iff_ord_connected ↔ convex.ord_connected _ diff --git a/src/analysis/convex/quasiconvex.lean b/src/analysis/convex/quasiconvex.lean index 74c304828e03f..351f7f7fd2222 100644 --- a/src/analysis/convex/quasiconvex.lean +++ b/src/analysis/convex/quasiconvex.lean @@ -111,10 +111,10 @@ lemma quasiconcave_on_iff_min_le : min (f x) (f y) ≤ f (a • x + b • y) := @quasiconvex_on_iff_le_max 𝕜 E βᵒᵈ _ _ _ _ _ _ -lemma quasilinear_on_iff_mem_interval : +lemma quasilinear_on_iff_mem_uIcc : quasilinear_on 𝕜 s f ↔ convex 𝕜 s ∧ ∀ ⦃x⦄, x ∈ s → ∀ ⦃y⦄, y ∈ s → ∀ ⦃a b : 𝕜⦄, 0 ≤ a → 0 ≤ b → a + b = 1 → - f (a • x + b • y) ∈ interval (f x) (f y) := + f (a • x + b • y) ∈ uIcc (f x) (f y) := begin rw [quasilinear_on, quasiconvex_on_iff_le_max, quasiconcave_on_iff_min_le, and_and_and_comm, and_self], @@ -198,7 +198,7 @@ variables [linear_ordered_field 𝕜] [linear_ordered_add_comm_monoid β] {s : s lemma quasilinear_on.monotone_on_or_antitone_on (hf : quasilinear_on 𝕜 s f) : monotone_on f s ∨ antitone_on f s := begin - simp_rw [monotone_on_or_antitone_on_iff_interval, ←segment_eq_interval], + simp_rw [monotone_on_or_antitone_on_iff_uIcc, ←segment_eq_uIcc], rintro a ha b hb c hc h, refine ⟨((hf.2 _).segment_subset _ _ h).2, ((hf.1 _).segment_subset _ _ h).2⟩; simp [*], end diff --git a/src/analysis/convex/segment.lean b/src/analysis/convex/segment.lean index ea56f0fa3b164..341637fafe168 100644 --- a/src/analysis/convex/segment.lean +++ b/src/analysis/convex/segment.lean @@ -379,22 +379,22 @@ end ordered_cancel_add_comm_monoid section linear_ordered_add_comm_monoid variables [linear_ordered_add_comm_monoid E] [module 𝕜 E] [ordered_smul 𝕜 E] {𝕜} {a b : 𝕜} -lemma segment_subset_interval (x y : E) : [x -[𝕜] y] ⊆ interval x y := +lemma segment_subset_uIcc (x y : E) : [x -[𝕜] y] ⊆ uIcc x y := begin cases le_total x y, - { rw interval_of_le h, + { rw uIcc_of_le h, exact segment_subset_Icc h }, - { rw [interval_of_ge h, segment_symm], + { rw [uIcc_of_ge h, segment_symm], exact segment_subset_Icc h } end lemma convex.min_le_combo (x y : E) (ha : 0 ≤ a) (hb : 0 ≤ b) (hab : a + b = 1) : min x y ≤ a • x + b • y := -(segment_subset_interval x y ⟨_, _, ha, hb, hab, rfl⟩).1 +(segment_subset_uIcc x y ⟨_, _, ha, hb, hab, rfl⟩).1 lemma convex.combo_le_max (x y : E) (ha : 0 ≤ a) (hb : 0 ≤ b) (hab : a + b = 1) : a • x + b • y ≤ max x y := -(segment_subset_interval x y ⟨_, _, ha, hb, hab, rfl⟩).2 +(segment_subset_uIcc x y ⟨_, _, ha, hb, hab, rfl⟩).2 end linear_ordered_add_comm_monoid end ordered_semiring @@ -440,7 +440,7 @@ begin { rw [open_segment_symm, open_segment_eq_Ioo h, max_eq_left h.le, min_eq_right h.le] } end -lemma segment_eq_interval (x y : 𝕜) : [x -[𝕜] y] = interval x y := segment_eq_Icc' _ _ +lemma segment_eq_uIcc (x y : 𝕜) : [x -[𝕜] y] = uIcc x y := segment_eq_Icc' _ _ /-- A point is in an `Icc` iff it can be expressed as a convex combination of the endpoints. -/ lemma convex.mem_Icc (h : x ≤ y) : diff --git a/src/analysis/convex/star.lean b/src/analysis/convex/star.lean index eab8812f833ba..e0a9eaeb1ec41 100644 --- a/src/analysis/convex/star.lean +++ b/src/analysis/convex/star.lean @@ -434,8 +434,7 @@ end lemma star_convex_iff_ord_connected [linear_ordered_field 𝕜] {x : 𝕜} {s : set 𝕜} (hx : x ∈ s) : star_convex 𝕜 x s ↔ s.ord_connected := -by simp_rw [ord_connected_iff_interval_subset_left hx, star_convex_iff_segment_subset, - segment_eq_interval] +by simp_rw [ord_connected_iff_uIcc_subset_left hx, star_convex_iff_segment_subset, segment_eq_uIcc] alias star_convex_iff_ord_connected ↔ star_convex.ord_connected _ diff --git a/src/analysis/convex/strict.lean b/src/analysis/convex/strict.lean index d2a226136a097..f444b82976233 100644 --- a/src/analysis/convex/strict.lean +++ b/src/analysis/convex/strict.lean @@ -159,10 +159,8 @@ lemma strict_convex_Icc (r s : β) : strict_convex 𝕜 (Icc r s) := ord_connect lemma strict_convex_Ioo (r s : β) : strict_convex 𝕜 (Ioo r s) := ord_connected_Ioo.strict_convex lemma strict_convex_Ico (r s : β) : strict_convex 𝕜 (Ico r s) := ord_connected_Ico.strict_convex lemma strict_convex_Ioc (r s : β) : strict_convex 𝕜 (Ioc r s) := ord_connected_Ioc.strict_convex -lemma strict_convex_interval (r s : β) : strict_convex 𝕜 (interval r s) := strict_convex_Icc _ _ - -lemma strict_convex_interval_oc (r s : β) : strict_convex 𝕜 (interval_oc r s) := -strict_convex_Ioc _ _ +lemma strict_convex_uIcc (r s : β) : strict_convex 𝕜 (uIcc r s) := strict_convex_Icc _ _ +lemma strict_convex_uIoc (r s : β) : strict_convex 𝕜 (uIoc r s) := strict_convex_Ioc _ _ end linear_ordered_cancel_add_comm_monoid end module diff --git a/src/analysis/special_functions/integrals.lean b/src/analysis/special_functions/integrals.lean index 1d09b92712e00..444560c44e200 100644 --- a/src/analysis/special_functions/integrals.lean +++ b/src/analysis/special_functions/integrals.lean @@ -60,7 +60,7 @@ begin { exact interval_integrable.trans (this a).symm (this b) }, have : ∀ (c : ℝ), (0 ≤ c) → interval_integrable (λ x, x ^ r) volume 0 c, { intros c hc, - rw [interval_integrable_iff, interval_oc_of_le hc], + rw [interval_integrable_iff, uIoc_of_le hc], have hderiv : ∀ x ∈ Ioo 0 c, has_deriv_at (λ x : ℝ, x ^ (r + 1) / (r + 1)) (x ^ r) x, { intros x hx, convert (real.has_deriv_at_rpow_const (or.inl hx.1.ne')).div_const (r + 1), field_simp [(by linarith : r + 1 ≠ 0)], ring, }, @@ -73,7 +73,7 @@ begin have m := (this (-c) (by linarith)).smul (cos (r * π)), rw interval_integrable_iff at m ⊢, refine m.congr_fun _ measurable_set_Ioc, intros x hx, - rw interval_oc_of_le (by linarith : 0 ≤ -c) at hx, + rw uIoc_of_le (by linarith : 0 ≤ -c) at hx, simp only [pi.smul_apply, algebra.id.smul_eq_mul, log_neg_eq_log, mul_comm, rpow_def_of_pos hx.1, rpow_def_of_neg (by linarith [hx.1] : -x < 0)], } end @@ -243,7 +243,7 @@ lemma integral_cpow {r : ℂ} (ha : 0 < a) (hb : 0 < b) (hr : r ≠ -1) : ∫ (x : ℝ) in a..b, (x : ℂ) ^ r = (b ^ (r + 1) - a ^ (r + 1)) / (r + 1) := begin rw sub_div, - suffices : ∀ x ∈ set.interval a b, has_deriv_at (λ z : ℂ, z ^ (r + 1) / (r + 1)) (x ^ r) x, + suffices : ∀ x ∈ set.uIcc a b, has_deriv_at (λ z : ℂ, z ^ (r + 1) / (r + 1)) (x ^ r) x, { exact integral_eq_sub_of_has_deriv_at (λ x hx, (this x hx).comp_of_real) (interval_integrable_cpow ha hb) }, intros x hx, @@ -269,27 +269,27 @@ by simpa using integral_zpow (or.inl (int.coe_nat_nonneg n)) /-- Integral of `|x - a| ^ n` over `Ι a b`. This integral appears in the proof of the Picard-Lindelöf/Cauchy-Lipschitz theorem. -/ -lemma integral_pow_abs_sub_interval_oc : +lemma integral_pow_abs_sub_uIoc : ∫ x in Ι a b, |x - a| ^ n = |b - a| ^ (n + 1) / (n + 1) := begin cases le_or_lt a b with hab hab, { calc ∫ x in Ι a b, |x - a| ^ n = ∫ x in a..b, |x - a| ^ n : - by rw [interval_oc_of_le hab, ← integral_of_le hab] + by rw [uIoc_of_le hab, ← integral_of_le hab] ... = ∫ x in 0..(b - a), x ^ n : begin simp only [integral_comp_sub_right (λ x, |x| ^ n), sub_self], refine integral_congr (λ x hx, congr_arg2 has_pow.pow (abs_of_nonneg $ _) rfl), - rw interval_of_le (sub_nonneg.2 hab) at hx, + rw uIcc_of_le (sub_nonneg.2 hab) at hx, exact hx.1 end ... = |b - a| ^ (n + 1) / (n + 1) : by simp [abs_of_nonneg (sub_nonneg.2 hab)] }, { calc ∫ x in Ι a b, |x - a| ^ n = ∫ x in b..a, |x - a| ^ n : - by rw [interval_oc_of_lt hab, ← integral_of_le hab.le] + by rw [uIoc_of_lt hab, ← integral_of_le hab.le] ... = ∫ x in b - a..0, (-x) ^ n : begin simp only [integral_comp_sub_right (λ x, |x| ^ n), sub_self], refine integral_congr (λ x hx, congr_arg2 has_pow.pow (abs_of_nonpos $ _) rfl), - rw interval_of_le (sub_nonpos.2 hab.le) at hx, + rw uIcc_of_le (sub_nonpos.2 hab.le) at hx, exact hx.2 end ... = |b - a| ^ (n + 1) / (n + 1) : @@ -313,16 +313,16 @@ begin have h' := λ x hx, ne_of_mem_of_not_mem hx h, rw [integral_deriv_eq_sub' _ deriv_log' (λ x hx, differentiable_at_log (h' x hx)) (continuous_on_inv₀.mono $ subset_compl_singleton_iff.mpr h), - log_div (h' b right_mem_interval) (h' a left_mem_interval)], + log_div (h' b right_mem_uIcc) (h' a left_mem_uIcc)], end @[simp] lemma integral_inv_of_pos (ha : 0 < a) (hb : 0 < b) : ∫ x in a..b, x⁻¹ = log (b / a) := -integral_inv $ not_mem_interval_of_lt ha hb +integral_inv $ not_mem_uIcc_of_lt ha hb @[simp] lemma integral_inv_of_neg (ha : a < 0) (hb : b < 0) : ∫ x in a..b, x⁻¹ = log (b / a) := -integral_inv $ not_mem_interval_of_gt ha hb +integral_inv $ not_mem_uIcc_of_gt ha hb lemma integral_one_div (h : (0:ℝ) ∉ [a, b]) : ∫ x : ℝ in a..b, 1/x = log (b / a) := by simp only [one_div, integral_inv h] @@ -365,12 +365,12 @@ end @[simp] lemma integral_log_of_pos (ha : 0 < a) (hb : 0 < b) : ∫ x in a..b, log x = b * log b - a * log a - b + a := -integral_log $ not_mem_interval_of_lt ha hb +integral_log $ not_mem_uIcc_of_lt ha hb @[simp] lemma integral_log_of_neg (ha : a < 0) (hb : b < 0) : ∫ x in a..b, log x = b * log b - a * log a - b + a := -integral_log $ not_mem_interval_of_gt ha hb +integral_log $ not_mem_uIcc_of_gt ha hb @[simp] lemma integral_sin : ∫ x in a..b, sin x = cos a - cos b := diff --git a/src/analysis/special_functions/non_integrable.lean b/src/analysis/special_functions/non_integrable.lean index 4fc6b9f5ab2bb..fcb9d043b7207 100644 --- a/src/analysis/special_functions/non_integrable.lean +++ b/src/analysis/special_functions/non_integrable.lean @@ -60,7 +60,7 @@ begin { rcases hfg.exists_nonneg with ⟨C, C₀, hC⟩, have h : ∀ᶠ x : ℝ × ℝ in l.prod l, ∀ y ∈ [x.1, x.2], (differentiable_at ℝ f y ∧ ‖deriv f y‖ ≤ C * ‖g y‖) ∧ y ∈ [a, b], - from (tendsto_fst.interval tendsto_snd).eventually ((hd.and hC.bound).and hl).small_sets, + from (tendsto_fst.uIcc tendsto_snd).eventually ((hd.and hC.bound).and hl).small_sets, rcases mem_prod_self_iff.1 h with ⟨s, hsl, hs⟩, simp only [prod_subset_iff, mem_set_of_eq] at hs, exact ⟨C, C₀, s, hsl, λ x hx y hy z hz, (hs x hx y hy z hz).2, @@ -74,9 +74,9 @@ begin specialize hsub c hc d hd, specialize hfd c hc d hd, replace hg : ∀ x ∈ Ι c d, ‖deriv f x‖ ≤ C * ‖g x‖, from λ z hz, hg c hc d hd z ⟨hz.1.le, hz.2⟩, have hg_ae : ∀ᵐ x ∂(volume.restrict (Ι c d)), ‖deriv f x‖ ≤ C * ‖g x‖, - from (ae_restrict_mem measurable_set_interval_oc).mono hg, + from (ae_restrict_mem measurable_set_uIoc).mono hg, have hsub' : Ι c d ⊆ Ι a b, - from interval_oc_subset_interval_oc_of_interval_subset_interval hsub, + from uIoc_subset_uIoc_of_uIcc_subset_uIcc hsub, have hfi : interval_integrable (deriv f) volume c d, from (hgi.mono_set hsub).mono_fun' (ae_strongly_measurable_deriv _ _) hg_ae, refine hlt.not_le (sub_le_iff_le_add'.1 _), @@ -85,7 +85,7 @@ begin ... = ‖∫ x in Ι c d, deriv f x‖ : norm_integral_eq_norm_integral_Ioc _ ... ≤ ∫ x in Ι c d, ‖deriv f x‖ : norm_integral_le_integral_norm _ ... ≤ ∫ x in Ι c d, C * ‖g x‖ : - set_integral_mono_on hfi.norm.def (hgi.def.mono_set hsub') measurable_set_interval_oc hg + set_integral_mono_on hfi.norm.def (hgi.def.mono_set hsub') measurable_set_uIoc hg ... ≤ ∫ x in Ι a b, C * ‖g x‖ : set_integral_mono_set hgi.def (ae_of_all _ $ λ x, mul_nonneg hC₀ (norm_nonneg _)) hsub'.eventually_le diff --git a/src/analysis/sum_integral_comparisons.lean b/src/analysis/sum_integral_comparisons.lean index f1c9759ecafa5..611b475c6490d 100644 --- a/src/analysis/sum_integral_comparisons.lean +++ b/src/analysis/sum_integral_comparisons.lean @@ -49,7 +49,7 @@ begin have hint : ∀ (k : ℕ), k < a → interval_integrable f volume (x₀+k) (x₀ + (k + 1 : ℕ)), { assume k hk, refine (hf.mono _).interval_integrable, - rw interval_of_le, + rw uIcc_of_le, { apply Icc_subset_Icc, { simp only [le_add_iff_nonneg_right, nat.cast_nonneg] }, { simp only [add_le_add_iff_left, nat.cast_le, nat.succ_le_of_lt hk] } }, @@ -91,7 +91,7 @@ begin have hint : ∀ (k : ℕ), k < a → interval_integrable f volume (x₀+k) (x₀ + (k + 1 : ℕ)), { assume k hk, refine (hf.mono _).interval_integrable, - rw interval_of_le, + rw uIcc_of_le, { apply Icc_subset_Icc, { simp only [le_add_iff_nonneg_right, nat.cast_nonneg] }, { simp only [add_le_add_iff_left, nat.cast_le, nat.succ_le_of_lt hk] } }, diff --git a/src/data/finset/locally_finite.lean b/src/data/finset/locally_finite.lean index c41563436e776..6e3cf7f748223 100644 --- a/src/data/finset/locally_finite.lean +++ b/src/data/finset/locally_finite.lean @@ -502,69 +502,62 @@ end linear_order section lattice variables [lattice α] [locally_finite_order α] {a a₁ a₂ b b₁ b₂ c x : α} -lemma interval_to_dual (a b : α) : [to_dual a, to_dual b] = [a, b].map to_dual.to_embedding := +lemma uIcc_to_dual (a b : α) : [to_dual a, to_dual b] = [a, b].map to_dual.to_embedding := Icc_to_dual _ _ -@[simp] lemma interval_of_le (h : a ≤ b) : [a, b] = Icc a b := -by rw [interval, inf_eq_left.2 h, sup_eq_right.2 h] +@[simp] lemma uIcc_of_le (h : a ≤ b) : [a, b] = Icc a b := +by rw [uIcc, inf_eq_left.2 h, sup_eq_right.2 h] -@[simp] lemma interval_of_ge (h : b ≤ a) : [a, b] = Icc b a := -by rw [interval, inf_eq_right.2 h, sup_eq_left.2 h] +@[simp] lemma uIcc_of_ge (h : b ≤ a) : [a, b] = Icc b a := +by rw [uIcc, inf_eq_right.2 h, sup_eq_left.2 h] -lemma interval_swap (a b : α) : [a, b] = [b, a] := by rw [interval, interval, inf_comm, sup_comm] +lemma uIcc_comm (a b : α) : [a, b] = [b, a] := by rw [uIcc, uIcc, inf_comm, sup_comm] -@[simp] lemma interval_self : [a, a] = {a} := by simp [interval] +@[simp] lemma uIcc_self : [a, a] = {a} := by simp [uIcc] -@[simp] lemma nonempty_interval : finset.nonempty [a, b] := nonempty_Icc.2 inf_le_sup +@[simp] lemma nonempty_uIcc : finset.nonempty [a, b] := nonempty_Icc.2 inf_le_sup -lemma Icc_subset_interval : Icc a b ⊆ [a, b] := Icc_subset_Icc inf_le_left le_sup_right -lemma Icc_subset_interval' : Icc b a ⊆ [a, b] := Icc_subset_Icc inf_le_right le_sup_left +lemma Icc_subset_uIcc : Icc a b ⊆ [a, b] := Icc_subset_Icc inf_le_left le_sup_right +lemma Icc_subset_uIcc' : Icc b a ⊆ [a, b] := Icc_subset_Icc inf_le_right le_sup_left -@[simp] lemma left_mem_interval : a ∈ [a, b] := mem_Icc.2 ⟨inf_le_left, le_sup_left⟩ -@[simp] lemma right_mem_interval : b ∈ [a, b] := mem_Icc.2 ⟨inf_le_right, le_sup_right⟩ +@[simp] lemma left_mem_uIcc : a ∈ [a, b] := mem_Icc.2 ⟨inf_le_left, le_sup_left⟩ +@[simp] lemma right_mem_uIcc : b ∈ [a, b] := mem_Icc.2 ⟨inf_le_right, le_sup_right⟩ -lemma mem_interval_of_le (ha : a ≤ x) (hb : x ≤ b) : x ∈ [a, b] := -Icc_subset_interval $ mem_Icc.2 ⟨ha, hb⟩ +lemma mem_uIcc_of_le (ha : a ≤ x) (hb : x ≤ b) : x ∈ [a, b] := Icc_subset_uIcc $ mem_Icc.2 ⟨ha, hb⟩ +lemma mem_uIcc_of_ge (hb : b ≤ x) (ha : x ≤ a) : x ∈ [a, b] := Icc_subset_uIcc' $ mem_Icc.2 ⟨hb, ha⟩ -lemma mem_interval_of_ge (hb : b ≤ x) (ha : x ≤ a) : x ∈ [a, b] := -Icc_subset_interval' $ mem_Icc.2 ⟨hb, ha⟩ +lemma uIcc_subset_uIcc (h₁ : a₁ ∈ [a₂, b₂]) (h₂ : b₁ ∈ [a₂, b₂]) : [a₁, b₁] ⊆ [a₂, b₂] := +by { rw mem_uIcc at h₁ h₂, exact Icc_subset_Icc (le_inf h₁.1 h₂.1) (sup_le h₁.2 h₂.2) } -lemma interval_subset_interval (h₁ : a₁ ∈ [a₂, b₂]) (h₂ : b₁ ∈ [a₂, b₂]) : [a₁, b₁] ⊆ [a₂, b₂] := -by { rw mem_interval at h₁ h₂, exact Icc_subset_Icc (le_inf h₁.1 h₂.1) (sup_le h₁.2 h₂.2) } - -lemma interval_subset_Icc (ha : a₁ ∈ Icc a₂ b₂) (hb : b₁ ∈ Icc a₂ b₂) : [a₁, b₁] ⊆ Icc a₂ b₂ := +lemma uIcc_subset_Icc (ha : a₁ ∈ Icc a₂ b₂) (hb : b₁ ∈ Icc a₂ b₂) : [a₁, b₁] ⊆ Icc a₂ b₂ := by { rw mem_Icc at ha hb, exact Icc_subset_Icc (le_inf ha.1 hb.1) (sup_le ha.2 hb.2) } -lemma interval_subset_interval_iff_mem : [a₁, b₁] ⊆ [a₂, b₂] ↔ a₁ ∈ [a₂, b₂] ∧ b₁ ∈ [a₂, b₂] := -⟨λ h, ⟨h left_mem_interval, h right_mem_interval⟩, λ h, interval_subset_interval h.1 h.2⟩ +lemma uIcc_subset_uIcc_iff_mem : [a₁, b₁] ⊆ [a₂, b₂] ↔ a₁ ∈ [a₂, b₂] ∧ b₁ ∈ [a₂, b₂] := +⟨λ h, ⟨h left_mem_uIcc, h right_mem_uIcc⟩, λ h, uIcc_subset_uIcc h.1 h.2⟩ -lemma interval_subset_interval_iff_le' : - [a₁, b₁] ⊆ [a₂, b₂] ↔ a₂ ⊓ b₂ ≤ a₁ ⊓ b₁ ∧ a₁ ⊔ b₁ ≤ a₂ ⊔ b₂ := +lemma uIcc_subset_uIcc_iff_le' : [a₁, b₁] ⊆ [a₂, b₂] ↔ a₂ ⊓ b₂ ≤ a₁ ⊓ b₁ ∧ a₁ ⊔ b₁ ≤ a₂ ⊔ b₂ := Icc_subset_Icc_iff inf_le_sup -lemma interval_subset_interval_right (h : x ∈ [a, b]) : [x, b] ⊆ [a, b] := -interval_subset_interval h right_mem_interval - -lemma interval_subset_interval_left (h : x ∈ [a, b]) : [a, x] ⊆ [a, b] := -interval_subset_interval left_mem_interval h +lemma uIcc_subset_uIcc_right (h : x ∈ [a, b]) : [x, b] ⊆ [a, b] := uIcc_subset_uIcc h right_mem_uIcc +lemma uIcc_subset_uIcc_left (h : x ∈ [a, b]) : [a, x] ⊆ [a, b] := uIcc_subset_uIcc left_mem_uIcc h end lattice section distrib_lattice variables [distrib_lattice α] [locally_finite_order α] {a a₁ a₂ b b₁ b₂ c x : α} -lemma eq_of_mem_interval_of_mem_interval : a ∈ [b, c] → b ∈ [a, c] → a = b := -by { simp_rw mem_interval, exact set.eq_of_mem_interval_of_mem_interval } +lemma eq_of_mem_uIcc_of_mem_uIcc : a ∈ [b, c] → b ∈ [a, c] → a = b := +by { simp_rw mem_uIcc, exact set.eq_of_mem_uIcc_of_mem_uIcc } -lemma eq_of_mem_interval_of_mem_interval' : b ∈ [a, c] → c ∈ [a, b] → b = c := -by { simp_rw mem_interval, exact set.eq_of_mem_interval_of_mem_interval' } +lemma eq_of_mem_uIcc_of_mem_uIcc' : b ∈ [a, c] → c ∈ [a, b] → b = c := +by { simp_rw mem_uIcc, exact set.eq_of_mem_uIcc_of_mem_uIcc' } -lemma interval_injective_right (a : α) : injective (λ b, [b, a]) := +lemma uIcc_injective_right (a : α) : injective (λ b, [b, a]) := λ b c h, by { rw ext_iff at h, - exact eq_of_mem_interval_of_mem_interval ((h _).1 left_mem_interval) ((h _).2 left_mem_interval) } + exact eq_of_mem_uIcc_of_mem_uIcc ((h _).1 left_mem_uIcc) ((h _).2 left_mem_uIcc) } -lemma interval_injective_left (a : α) : injective (interval a) := -by simpa only [interval_swap] using interval_injective_right a +lemma uIcc_injective_left (a : α) : injective (uIcc a) := +by simpa only [uIcc_comm] using uIcc_injective_right a end distrib_lattice @@ -573,27 +566,27 @@ variables [linear_order α] [locally_finite_order α] {a a₁ a₂ b b₁ b₂ c lemma Icc_min_max : Icc (min a b) (max a b) = [a, b] := rfl -lemma interval_of_not_le (h : ¬ a ≤ b) : [a, b] = Icc b a := interval_of_ge $ le_of_not_ge h -lemma interval_of_not_ge (h : ¬ b ≤ a) : [a, b] = Icc a b := interval_of_le $ le_of_not_ge h +lemma uIcc_of_not_le (h : ¬ a ≤ b) : [a, b] = Icc b a := uIcc_of_ge $ le_of_not_ge h +lemma uIcc_of_not_ge (h : ¬ b ≤ a) : [a, b] = Icc a b := uIcc_of_le $ le_of_not_ge h -lemma interval_eq_union : [a, b] = Icc a b ∪ Icc b a := -coe_injective $ by { push_cast, exact set.interval_eq_union } +lemma uIcc_eq_union : [a, b] = Icc a b ∪ Icc b a := +coe_injective $ by { push_cast, exact set.uIcc_eq_union } -lemma mem_interval' : a ∈ [b, c] ↔ b ≤ a ∧ a ≤ c ∨ c ≤ a ∧ a ≤ b := by simp [interval_eq_union] +lemma mem_uIcc' : a ∈ [b, c] ↔ b ≤ a ∧ a ≤ c ∨ c ≤ a ∧ a ≤ b := by simp [uIcc_eq_union] -lemma not_mem_interval_of_lt : c < a → c < b → c ∉ [a, b] := -by { rw mem_interval, exact set.not_mem_interval_of_lt } +lemma not_mem_uIcc_of_lt : c < a → c < b → c ∉ [a, b] := +by { rw mem_uIcc, exact set.not_mem_uIcc_of_lt } -lemma not_mem_interval_of_gt : a < c → b < c → c ∉ [a, b] := -by { rw mem_interval, exact set.not_mem_interval_of_gt } +lemma not_mem_uIcc_of_gt : a < c → b < c → c ∉ [a, b] := +by { rw mem_uIcc, exact set.not_mem_uIcc_of_gt } -lemma interval_subset_interval_iff_le : +lemma uIcc_subset_uIcc_iff_le : [a₁, b₁] ⊆ [a₂, b₂] ↔ min a₂ b₂ ≤ min a₁ b₁ ∧ max a₁ b₁ ≤ max a₂ b₂ := -interval_subset_interval_iff_le' +uIcc_subset_uIcc_iff_le' /-- A sort of triangle inequality. -/ -lemma interval_subset_interval_union_interval : [a, c] ⊆ [a, b] ∪ [b, c] := -coe_subset.1 $ by { push_cast, exact set.interval_subset_interval_union_interval } +lemma uIcc_subset_uIcc_union_uIcc : [a, c] ⊆ [a, b] ∪ [b, c] := +coe_subset.1 $ by { push_cast, exact set.uIcc_subset_uIcc_union_uIcc } end linear_order diff --git a/src/data/set/intervals/ord_connected.lean b/src/data/set/intervals/ord_connected.lean index 50952b855241b..7dd30593bff4a 100644 --- a/src/data/set/intervals/ord_connected.lean +++ b/src/data/set/intervals/ord_connected.lean @@ -176,35 +176,35 @@ end preorder section linear_order variables {α : Type*} [linear_order α] {s : set α} {x : α} -@[instance] lemma ord_connected_interval {a b : α} : ord_connected [a, b] := ord_connected_Icc -@[instance] lemma ord_connected_interval_oc {a b : α} : ord_connected (Ι a b) := ord_connected_Ioc +@[instance] lemma ord_connected_uIcc {a b : α} : ord_connected [a, b] := ord_connected_Icc +@[instance] lemma ord_connected_uIoc {a b : α} : ord_connected (Ι a b) := ord_connected_Ioc -lemma ord_connected.interval_subset (hs : ord_connected s) ⦃x⦄ (hx : x ∈ s) ⦃y⦄ (hy : y ∈ s) : +lemma ord_connected.uIcc_subset (hs : ord_connected s) ⦃x⦄ (hx : x ∈ s) ⦃y⦄ (hy : y ∈ s) : [x, y] ⊆ s := hs.out (min_rec' (∈ s) hx hy) (max_rec' (∈ s) hx hy) -lemma ord_connected.interval_oc_subset (hs : ord_connected s) ⦃x⦄ (hx : x ∈ s) ⦃y⦄ (hy : y ∈ s) : +lemma ord_connected.uIoc_subset (hs : ord_connected s) ⦃x⦄ (hx : x ∈ s) ⦃y⦄ (hy : y ∈ s) : Ι x y ⊆ s := -Ioc_subset_Icc_self.trans $ hs.interval_subset hx hy +Ioc_subset_Icc_self.trans $ hs.uIcc_subset hx hy -lemma ord_connected_iff_interval_subset : +lemma ord_connected_iff_uIcc_subset : ord_connected s ↔ ∀ ⦃x⦄ (hx : x ∈ s) ⦃y⦄ (hy : y ∈ s), [x, y] ⊆ s := -⟨λ h, h.interval_subset, λ H, ⟨λ x hx y hy, Icc_subset_interval.trans $ H hx hy⟩⟩ +⟨λ h, h.uIcc_subset, λ H, ⟨λ x hx y hy, Icc_subset_uIcc.trans $ H hx hy⟩⟩ -lemma ord_connected_of_interval_subset_left (h : ∀ y ∈ s, [x, y] ⊆ s) : +lemma ord_connected_of_uIcc_subset_left (h : ∀ y ∈ s, [x, y] ⊆ s) : ord_connected s := -ord_connected_iff_interval_subset.2 $ λ y hy z hz, -calc [y, z] ⊆ [y, x] ∪ [x, z] : interval_subset_interval_union_interval -... = [x, y] ∪ [x, z] : by rw [interval_swap] +ord_connected_iff_uIcc_subset.2 $ λ y hy z hz, +calc [y, z] ⊆ [y, x] ∪ [x, z] : uIcc_subset_uIcc_union_uIcc +... = [x, y] ∪ [x, z] : by rw [uIcc_comm] ... ⊆ s : union_subset (h y hy) (h z hz) -lemma ord_connected_iff_interval_subset_left (hx : x ∈ s) : +lemma ord_connected_iff_uIcc_subset_left (hx : x ∈ s) : ord_connected s ↔ ∀ ⦃y⦄, y ∈ s → [x, y] ⊆ s := -⟨λ hs, hs.interval_subset hx, ord_connected_of_interval_subset_left⟩ +⟨λ hs, hs.uIcc_subset hx, ord_connected_of_uIcc_subset_left⟩ -lemma ord_connected_iff_interval_subset_right (hx : x ∈ s) : +lemma ord_connected_iff_uIcc_subset_right (hx : x ∈ s) : ord_connected s ↔ ∀ ⦃y⦄, y ∈ s → [y, x] ⊆ s := -by simp_rw [ord_connected_iff_interval_subset_left hx, interval_swap] +by simp_rw [ord_connected_iff_uIcc_subset_left hx, uIcc_comm] end linear_order end set diff --git a/src/data/set/intervals/ord_connected_component.lean b/src/data/set/intervals/ord_connected_component.lean index 3bc77087f89db..eff4c832d07d6 100644 --- a/src/data/set/intervals/ord_connected_component.lean +++ b/src/data/set/intervals/ord_connected_component.lean @@ -13,7 +13,7 @@ import tactic.wlog > Any changes to this file require a corresponding PR to mathlib4. In this file we define `set.ord_connected_component s x` to be the set of `y` such that -`set.interval x y ⊆ s` and prove some basic facts about this definition. At the moment of writing, +`set.uIcc x y ⊆ s` and prove some basic facts about this definition. At the moment of writing, this construction is used only to prove that any linear order with order topology is a T₅ space, so we only add API needed for this lemma. -/ @@ -26,7 +26,7 @@ namespace set variables {α : Type*} [linear_order α] {s t : set α} {x y z : α} /-- Order-connected component of a point `x` in a set `s`. It is defined as the set of `y` such that -`set.interval x y ⊆ s`. Note that it is empty if and only if `x ∉ s`. -/ +`set.uIcc x y ⊆ s`. Note that it is empty if and only if `x ∉ s`. -/ def ord_connected_component (s : set α) (x : α) : set α := {y | [x, y] ⊆ s} lemma mem_ord_connected_component : y ∈ ord_connected_component s x ↔ [x, y] ⊆ s := iff.rfl @@ -34,20 +34,19 @@ lemma mem_ord_connected_component : y ∈ ord_connected_component s x ↔ [x, y] lemma dual_ord_connected_component : ord_connected_component (of_dual ⁻¹' s) (to_dual x) = of_dual ⁻¹' (ord_connected_component s x) := ext $ to_dual.surjective.forall.2 $ λ x, - by { rw [mem_ord_connected_component, dual_interval], refl } + by { rw [mem_ord_connected_component, dual_uIcc], refl } -lemma ord_connected_component_subset : ord_connected_component s x ⊆ s := -λ y hy, hy right_mem_interval +lemma ord_connected_component_subset : ord_connected_component s x ⊆ s := λ y hy, hy right_mem_uIcc lemma subset_ord_connected_component {t} [h : ord_connected s] (hs : x ∈ s) (ht : s ⊆ t) : s ⊆ ord_connected_component t x := -λ y hy, (h.interval_subset hs hy).trans ht +λ y hy, (h.uIcc_subset hs hy).trans ht @[simp] lemma self_mem_ord_connected_component : x ∈ ord_connected_component s x ↔ x ∈ s := -by rw [mem_ord_connected_component, interval_self, singleton_subset_iff] +by rw [mem_ord_connected_component, uIcc_self, singleton_subset_iff] @[simp] lemma nonempty_ord_connected_component : (ord_connected_component s x).nonempty ↔ x ∈ s := -⟨λ ⟨y, hy⟩, hy $ left_mem_interval, λ h, ⟨x, self_mem_ord_connected_component.2 h⟩⟩ +⟨λ ⟨y, hy⟩, hy $ left_mem_uIcc, λ h, ⟨x, self_mem_ord_connected_component.2 h⟩⟩ @[simp] lemma ord_connected_component_eq_empty : ord_connected_component s x = ∅ ↔ x ∉ s := by rw [← not_nonempty_iff_eq_empty, nonempty_ord_connected_component] @@ -64,11 +63,11 @@ by simp [ord_connected_component, set_of_and] lemma mem_ord_connected_component_comm : y ∈ ord_connected_component s x ↔ x ∈ ord_connected_component s y := -by rw [mem_ord_connected_component, mem_ord_connected_component, interval_swap] +by rw [mem_ord_connected_component, mem_ord_connected_component, uIcc_comm] lemma mem_ord_connected_component_trans (hxy : y ∈ ord_connected_component s x) (hyz : z ∈ ord_connected_component s y) : z ∈ ord_connected_component s x := -calc [x, z] ⊆ [x, y] ∪ [y, z] : interval_subset_interval_union_interval +calc [x, z] ⊆ [x, y] ∪ [y, z] : uIcc_subset_uIcc_union_uIcc ... ⊆ s : union_subset hxy hyz lemma ord_connected_component_eq (h : [x, y] ⊆ s) : @@ -77,7 +76,7 @@ ext $ λ z, ⟨mem_ord_connected_component_trans (mem_ord_connected_component_co mem_ord_connected_component_trans h⟩ instance : ord_connected (ord_connected_component s x) := -ord_connected_of_interval_subset_left $ λ y hy z hz, (interval_subset_interval_left hz).trans hy +ord_connected_of_uIcc_subset_left $ λ y hy z hz, (uIcc_subset_uIcc_left hz).trans hy /-- Projection from `s : set α` to `α` sending each order connected component of `s` to a single point of this component. -/ @@ -123,7 +122,7 @@ end lemma ord_connected_section_subset : ord_connected_section s ⊆ s := range_subset_iff.2 $ λ x, ord_connected_component_subset $ nonempty.some_mem _ -lemma eq_of_mem_ord_connected_section_of_interval_subset (hx : x ∈ ord_connected_section s) +lemma eq_of_mem_ord_connected_section_of_uIcc_subset (hx : x ∈ ord_connected_section s) (hy : y ∈ ord_connected_section s) (h : [x, y] ⊆ s) : x = y := begin rcases hx with ⟨x, rfl⟩, rcases hy with ⟨y, rfl⟩, @@ -170,25 +169,25 @@ begin rotate, from λ h₁ h₂ h₃ h₄, this h₂ h₁ h₄ h₃, cases ha with ha ha', cases hb with hb hb', have hsub : [a, b] ⊆ (ord_separating_set s t).ord_connected_sectionᶜ, - { rw [ord_separating_set_comm, interval_swap] at hb', - calc [a, b] ⊆ [a, x] ∪ [x, b] : interval_subset_interval_union_interval + { rw [ord_separating_set_comm, uIcc_comm] at hb', + calc [a, b] ⊆ [a, x] ∪ [x, b] : uIcc_subset_uIcc_union_uIcc ... ⊆ (ord_separating_set s t).ord_connected_sectionᶜ : union_subset ha' hb' }, clear ha' hb', cases le_total x a with hxa hax, - { exact hb (Icc_subset_interval' ⟨hxa, hab⟩) has }, + { exact hb (Icc_subset_uIcc' ⟨hxa, hab⟩) has }, cases le_total b x with hbx hxb, - { exact ha (Icc_subset_interval ⟨hab, hbx⟩) hbt }, + { exact ha (Icc_subset_uIcc ⟨hab, hbx⟩) hbt }, have : x ∈ ord_separating_set s t, { exact ⟨mem_Union₂.2 ⟨a, has, ha⟩, mem_Union₂.2 ⟨b, hbt, hb⟩⟩ }, lift x to ord_separating_set s t using this, suffices : ord_connected_component (ord_separating_set s t) x ⊆ [a, b], from hsub (this $ ord_connected_proj_mem_ord_connected_component _ _) (mem_range_self _), rintros y (hy : [↑x, y] ⊆ ord_separating_set s t), - rw [interval_of_le hab, mem_Icc, ← not_lt, ← not_lt], + rw [uIcc_of_le hab, mem_Icc, ← not_lt, ← not_lt], exact ⟨λ hya, disjoint_left.1 disjoint_left_ord_separating_set has - (hy $ Icc_subset_interval' ⟨hya.le, hax⟩), + (hy $ Icc_subset_uIcc' ⟨hya.le, hax⟩), λ hyb, disjoint_left.1 disjoint_right_ord_separating_set hbt - (hy $ Icc_subset_interval ⟨hxb, hyb.le⟩)⟩ + (hy $ Icc_subset_uIcc ⟨hxb, hyb.le⟩)⟩ end end set diff --git a/src/data/set/intervals/unordered_interval.lean b/src/data/set/intervals/unordered_interval.lean index 7dd4efbc72cc0..2bd90389e2153 100644 --- a/src/data/set/intervals/unordered_interval.lean +++ b/src/data/set/intervals/unordered_interval.lean @@ -12,25 +12,25 @@ import data.set.intervals.basic > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. > Any changes to this file require a corresponding PR to mathlib4. -In any lattice `α`, we define `interval a b` to be `Icc (a ⊓ b) (a ⊔ b)`, which in a linear order is -the set of elements lying between `a` and `b`. +In any lattice `α`, we define `uIcc a b` to be `Icc (a ⊓ b) (a ⊔ b)`, which in a linear order is the +set of elements lying between `a` and `b`. `Icc a b` requires the assumption `a ≤ b` to be meaningful, which is sometimes inconvenient. The interval as defined in this file is always the set of things lying between `a` and `b`, regardless of the relative order of `a` and `b`. -For real numbers, `interval a b` is the same as `segment ℝ a b`. +For real numbers, `uIcc a b` is the same as `segment ℝ a b`. -In a product or pi type, `interval a b` is the smallest box containing `a` and `b`. For example, -`interval (1, -1) (-1, 1) = Icc (-1, -1) (1, 1)` is the square of vertices `(1, -1)`, `(-1, -1)`, +In a product or pi type, `uIcc a b` is the smallest box containing `a` and `b`. For example, +`uIcc (1, -1) (-1, 1) = Icc (-1, -1) (1, 1)` is the square of vertices `(1, -1)`, `(-1, -1)`, `(-1, 1)`, `(1, 1)`. -In `finset α` (seen as a hypercube of dimension `fintype.card α`), `interval a b` is the smallest +In `finset α` (seen as a hypercube of dimension `fintype.card α`), `uIcc a b` is the smallest subcube containing both `a` and `b`. ## Notation -We use the localized notation `[a, b]` for `interval a b`. One can open the locale `interval` to +We use the localized notation `[a, b]` for `uIcc a b`. One can open the locale `interval` to make the notation available. -/ @@ -43,65 +43,58 @@ namespace set section lattice variables [lattice α] {a a₁ a₂ b b₁ b₂ c x : α} -/-- `interval a b` is the set of elements lying between `a` and `b`, with `a` and `b` included. +/-- `uIcc a b` is the set of elements lying between `a` and `b`, with `a` and `b` included. Note that we define it more generally in a lattice as `set.Icc (a ⊓ b) (a ⊔ b)`. In a product type, -`interval` corresponds to the bounding box of the two elements. -/ -def interval (a b : α) : set α := Icc (a ⊓ b) (a ⊔ b) +`uIcc` corresponds to the bounding box of the two elements. -/ +def uIcc (a b : α) : set α := Icc (a ⊓ b) (a ⊔ b) -localized "notation (name := set.interval) `[`a `, ` b `]` := set.interval a b" in interval +localized "notation (name := set.uIcc) `[`a `, ` b `]` := set.uIcc a b" in interval -@[simp] lemma dual_interval (a b : α) : [to_dual a, to_dual b] = of_dual ⁻¹' [a, b] := dual_Icc +@[simp] lemma dual_uIcc (a b : α) : [to_dual a, to_dual b] = of_dual ⁻¹' [a, b] := dual_Icc -@[simp] lemma interval_of_le (h : a ≤ b) : [a, b] = Icc a b := -by rw [interval, inf_eq_left.2 h, sup_eq_right.2 h] +@[simp] lemma uIcc_of_le (h : a ≤ b) : [a, b] = Icc a b := +by rw [uIcc, inf_eq_left.2 h, sup_eq_right.2 h] -@[simp] lemma interval_of_ge (h : b ≤ a) : [a, b] = Icc b a := -by rw [interval, inf_eq_right.2 h, sup_eq_left.2 h] +@[simp] lemma uIcc_of_ge (h : b ≤ a) : [a, b] = Icc b a := +by rw [uIcc, inf_eq_right.2 h, sup_eq_left.2 h] -lemma interval_swap (a b : α) : [a, b] = [b, a] := by simp_rw [interval, inf_comm, sup_comm] +lemma uIcc_comm (a b : α) : [a, b] = [b, a] := by simp_rw [uIcc, inf_comm, sup_comm] -lemma interval_of_lt (h : a < b) : [a, b] = Icc a b := -interval_of_le (le_of_lt h) +lemma uIcc_of_lt (h : a < b) : [a, b] = Icc a b := uIcc_of_le h.le +lemma uIcc_of_gt (h : b < a) : [a, b] = Icc b a := uIcc_of_ge h.le -lemma interval_of_gt (h : b < a) : [a, b] = Icc b a := -interval_of_ge (le_of_lt h) +@[simp] lemma uIcc_self : [a, a] = {a} := by simp [uIcc] -@[simp] lemma interval_self : [a, a] = {a} := by simp [interval] +@[simp] lemma nonempty_uIcc : [a, b].nonempty := nonempty_Icc.2 inf_le_sup -@[simp] lemma nonempty_interval : [a, b].nonempty := nonempty_Icc.2 inf_le_sup +lemma Icc_subset_uIcc : Icc a b ⊆ [a, b] := Icc_subset_Icc inf_le_left le_sup_right +lemma Icc_subset_uIcc' : Icc b a ⊆ [a, b] := Icc_subset_Icc inf_le_right le_sup_left -lemma Icc_subset_interval : Icc a b ⊆ [a, b] := Icc_subset_Icc inf_le_left le_sup_right -lemma Icc_subset_interval' : Icc b a ⊆ [a, b] := Icc_subset_Icc inf_le_right le_sup_left +@[simp] lemma left_mem_uIcc : a ∈ [a, b] := ⟨inf_le_left, le_sup_left⟩ +@[simp] lemma right_mem_uIcc : b ∈ [a, b] := ⟨inf_le_right, le_sup_right⟩ -@[simp] lemma left_mem_interval : a ∈ [a, b] := ⟨inf_le_left, le_sup_left⟩ -@[simp] lemma right_mem_interval : b ∈ [a, b] := ⟨inf_le_right, le_sup_right⟩ +lemma mem_uIcc_of_le (ha : a ≤ x) (hb : x ≤ b) : x ∈ [a, b] := Icc_subset_uIcc ⟨ha, hb⟩ +lemma mem_uIcc_of_ge (hb : b ≤ x) (ha : x ≤ a) : x ∈ [a, b] := Icc_subset_uIcc' ⟨hb, ha⟩ -lemma mem_interval_of_le (ha : a ≤ x) (hb : x ≤ b) : x ∈ [a, b] := Icc_subset_interval ⟨ha, hb⟩ -lemma mem_interval_of_ge (hb : b ≤ x) (ha : x ≤ a) : x ∈ [a, b] := Icc_subset_interval' ⟨hb, ha⟩ - -lemma interval_subset_interval (h₁ : a₁ ∈ [a₂, b₂]) (h₂ : b₁ ∈ [a₂, b₂]) : [a₁, b₁] ⊆ [a₂, b₂] := +lemma uIcc_subset_uIcc (h₁ : a₁ ∈ [a₂, b₂]) (h₂ : b₁ ∈ [a₂, b₂]) : [a₁, b₁] ⊆ [a₂, b₂] := Icc_subset_Icc (le_inf h₁.1 h₂.1) (sup_le h₁.2 h₂.2) -lemma interval_subset_Icc (ha : a₁ ∈ Icc a₂ b₂) (hb : b₁ ∈ Icc a₂ b₂) : [a₁, b₁] ⊆ Icc a₂ b₂ := +lemma uIcc_subset_Icc (ha : a₁ ∈ Icc a₂ b₂) (hb : b₁ ∈ Icc a₂ b₂) : [a₁, b₁] ⊆ Icc a₂ b₂ := Icc_subset_Icc (le_inf ha.1 hb.1) (sup_le ha.2 hb.2) -lemma interval_subset_interval_iff_mem : [a₁, b₁] ⊆ [a₂, b₂] ↔ a₁ ∈ [a₂, b₂] ∧ b₁ ∈ [a₂, b₂] := -iff.intro (λh, ⟨h left_mem_interval, h right_mem_interval⟩) (λ h, interval_subset_interval h.1 h.2) +lemma uIcc_subset_uIcc_iff_mem : [a₁, b₁] ⊆ [a₂, b₂] ↔ a₁ ∈ [a₂, b₂] ∧ b₁ ∈ [a₂, b₂] := +iff.intro (λh, ⟨h left_mem_uIcc, h right_mem_uIcc⟩) (λ h, uIcc_subset_uIcc h.1 h.2) -lemma interval_subset_interval_iff_le' : - [a₁, b₁] ⊆ [a₂, b₂] ↔ a₂ ⊓ b₂ ≤ a₁ ⊓ b₁ ∧ a₁ ⊔ b₁ ≤ a₂ ⊔ b₂ := +lemma uIcc_subset_uIcc_iff_le' : [a₁, b₁] ⊆ [a₂, b₂] ↔ a₂ ⊓ b₂ ≤ a₁ ⊓ b₁ ∧ a₁ ⊔ b₁ ≤ a₂ ⊔ b₂ := Icc_subset_Icc_iff inf_le_sup -lemma interval_subset_interval_right (h : x ∈ [a, b]) : [x, b] ⊆ [a, b] := -interval_subset_interval h right_mem_interval - -lemma interval_subset_interval_left (h : x ∈ [a, b]) : [a, x] ⊆ [a, b] := -interval_subset_interval left_mem_interval h +lemma uIcc_subset_uIcc_right (h : x ∈ [a, b]) : [x, b] ⊆ [a, b] := uIcc_subset_uIcc h right_mem_uIcc +lemma uIcc_subset_uIcc_left (h : x ∈ [a, b]) : [a, x] ⊆ [a, b] := uIcc_subset_uIcc left_mem_uIcc h -lemma bdd_below_bdd_above_iff_subset_interval (s : set α) : +lemma bdd_below_bdd_above_iff_subset_uIcc (s : set α) : bdd_below s ∧ bdd_above s ↔ ∃ a b, s ⊆ [a, b] := bdd_below_bdd_above_iff_subset_Icc.trans - ⟨λ ⟨a, b, h⟩, ⟨a, b, λ x hx, Icc_subset_interval (h hx)⟩, λ ⟨a, b, h⟩, ⟨_, _, h⟩⟩ + ⟨λ ⟨a, b, h⟩, ⟨a, b, λ x hx, Icc_subset_uIcc (h hx)⟩, λ ⟨a, b, h⟩, ⟨_, _, h⟩⟩ end lattice @@ -110,49 +103,48 @@ open_locale interval section distrib_lattice variables [distrib_lattice α] {a a₁ a₂ b b₁ b₂ c x : α} -lemma eq_of_mem_interval_of_mem_interval (ha : a ∈ [b, c]) (hb : b ∈ [a, c]) : a = b := +lemma eq_of_mem_uIcc_of_mem_uIcc (ha : a ∈ [b, c]) (hb : b ∈ [a, c]) : a = b := eq_of_inf_eq_sup_eq (inf_congr_right ha.1 hb.1) $ sup_congr_right ha.2 hb.2 -lemma eq_of_mem_interval_of_mem_interval' : b ∈ [a, c] → c ∈ [a, b] → b = c := -by simpa only [interval_swap a] using eq_of_mem_interval_of_mem_interval +lemma eq_of_mem_uIcc_of_mem_uIcc' : b ∈ [a, c] → c ∈ [a, b] → b = c := +by simpa only [uIcc_comm a] using eq_of_mem_uIcc_of_mem_uIcc -lemma interval_injective_right (a : α) : injective (λ b, interval b a) := +lemma uIcc_injective_right (a : α) : injective (λ b, uIcc b a) := λ b c h, by { rw ext_iff at h, - exact eq_of_mem_interval_of_mem_interval ((h _).1 left_mem_interval) ((h _).2 left_mem_interval) } + exact eq_of_mem_uIcc_of_mem_uIcc ((h _).1 left_mem_uIcc) ((h _).2 left_mem_uIcc) } -lemma interval_injective_left (a : α) : injective (interval a) := -by simpa only [interval_swap] using interval_injective_right a +lemma uIcc_injective_left (a : α) : injective (uIcc a) := +by simpa only [uIcc_comm] using uIcc_injective_right a end distrib_lattice section linear_order -variables [linear_order α] [linear_order β] {f : α → β} {s : set α} - {a a₁ a₂ b b₁ b₂ c x : α} +variables [linear_order α] [linear_order β] {f : α → β} {s : set α} {a a₁ a₂ b b₁ b₂ c d x : α} lemma Icc_min_max : Icc (min a b) (max a b) = [a, b] := rfl -lemma interval_of_not_le (h : ¬ a ≤ b) : [a, b] = Icc b a := interval_of_gt $ lt_of_not_ge h -lemma interval_of_not_ge (h : ¬ b ≤ a) : [a, b] = Icc a b := interval_of_lt $ lt_of_not_ge h +lemma uIcc_of_not_le (h : ¬ a ≤ b) : [a, b] = Icc b a := uIcc_of_gt $ lt_of_not_ge h +lemma uIcc_of_not_ge (h : ¬ b ≤ a) : [a, b] = Icc a b := uIcc_of_lt $ lt_of_not_ge h -lemma interval_eq_union : [a, b] = Icc a b ∪ Icc b a := by rw [Icc_union_Icc', max_comm]; refl +lemma uIcc_eq_union : [a, b] = Icc a b ∪ Icc b a := by rw [Icc_union_Icc', max_comm]; refl -lemma mem_interval : a ∈ [b, c] ↔ b ≤ a ∧ a ≤ c ∨ c ≤ a ∧ a ≤ b := by simp [interval_eq_union] +lemma mem_uIcc : a ∈ [b, c] ↔ b ≤ a ∧ a ≤ c ∨ c ≤ a ∧ a ≤ b := by simp [uIcc_eq_union] -lemma not_mem_interval_of_lt (ha : c < a) (hb : c < b) : c ∉ [a, b] := +lemma not_mem_uIcc_of_lt (ha : c < a) (hb : c < b) : c ∉ [a, b] := not_mem_Icc_of_lt $ lt_min_iff.mpr ⟨ha, hb⟩ -lemma not_mem_interval_of_gt (ha : a < c) (hb : b < c) : c ∉ [a, b] := +lemma not_mem_uIcc_of_gt (ha : a < c) (hb : b < c) : c ∉ [a, b] := not_mem_Icc_of_gt $ max_lt_iff.mpr ⟨ha, hb⟩ -lemma interval_subset_interval_iff_le : +lemma uIcc_subset_uIcc_iff_le : [a₁, b₁] ⊆ [a₂, b₂] ↔ min a₂ b₂ ≤ min a₁ b₁ ∧ max a₁ b₁ ≤ max a₂ b₂ := -interval_subset_interval_iff_le' +uIcc_subset_uIcc_iff_le' /-- A sort of triangle inequality. -/ -lemma interval_subset_interval_union_interval : [a, c] ⊆ [a, b] ∪ [b, c] := -λ x, by simp only [mem_interval, mem_union]; cases le_total a c; cases le_total x b; tauto +lemma uIcc_subset_uIcc_union_uIcc : [a, c] ⊆ [a, b] ∪ [b, c] := +λ x, by simp only [mem_uIcc, mem_union]; cases le_total a c; cases le_total x b; tauto -lemma monotone_or_antitone_iff_interval : +lemma monotone_or_antitone_iff_uIcc : monotone f ∨ antitone f ↔ ∀ a b c, c ∈ [a, b] → f c ∈ [f a, f b] := begin split, @@ -161,85 +153,76 @@ begin contrapose!, rw not_monotone_not_antitone_iff_exists_le_le, rintro ⟨a, b, c, hab, hbc, ⟨hfab, hfcb⟩ | ⟨hfba, hfbc⟩⟩, - { exact ⟨a, c, b, Icc_subset_interval ⟨hab, hbc⟩, λ h, h.2.not_lt $ max_lt hfab hfcb⟩ }, - { exact ⟨a, c, b, Icc_subset_interval ⟨hab, hbc⟩, λ h, h.1.not_lt $ lt_min hfba hfbc⟩ } + { exact ⟨a, c, b, Icc_subset_uIcc ⟨hab, hbc⟩, λ h, h.2.not_lt $ max_lt hfab hfcb⟩ }, + { exact ⟨a, c, b, Icc_subset_uIcc ⟨hab, hbc⟩, λ h, h.1.not_lt $ lt_min hfba hfbc⟩ } end -lemma monotone_on_or_antitone_on_iff_interval : +lemma monotone_on_or_antitone_on_iff_uIcc : monotone_on f s ∨ antitone_on f s ↔ ∀ a b c ∈ s, c ∈ [a, b] → f c ∈ [f a, f b] := -by simp [monotone_on_iff_monotone, antitone_on_iff_antitone, monotone_or_antitone_iff_interval, - mem_interval] +by simp [monotone_on_iff_monotone, antitone_on_iff_antitone, monotone_or_antitone_iff_uIcc, + mem_uIcc] /-- The open-closed interval with unordered bounds. -/ -def interval_oc : α → α → set α := λ a b, Ioc (min a b) (max a b) +def uIoc : α → α → set α := λ a b, Ioc (min a b) (max a b) -- Below is a capital iota -localized "notation `Ι` := set.interval_oc" in interval +localized "notation `Ι` := set.uIoc" in interval -@[simp] lemma interval_oc_of_le (h : a ≤ b) : Ι a b = Ioc a b := -by simp [interval_oc, h] +@[simp] lemma uIoc_of_le (h : a ≤ b) : Ι a b = Ioc a b := by simp [uIoc, h] +@[simp] lemma uIoc_of_lt (h : b < a) : Ι a b = Ioc b a := by simp [uIoc, h.le] -@[simp] lemma interval_oc_of_lt (h : b < a) : Ι a b = Ioc b a := -by simp [interval_oc, le_of_lt h] +lemma uIoc_eq_union : Ι a b = Ioc a b ∪ Ioc b a := by cases le_total a b; simp [uIoc, *] -lemma interval_oc_eq_union : Ι a b = Ioc a b ∪ Ioc b a := -by cases le_total a b; simp [interval_oc, *] +lemma mem_uIoc : a ∈ Ι b c ↔ b < a ∧ a ≤ c ∨ c < a ∧ a ≤ b := +by simp only [uIoc_eq_union, mem_union, mem_Ioc] -lemma mem_interval_oc : a ∈ Ι b c ↔ b < a ∧ a ≤ c ∨ c < a ∧ a ≤ b := -by simp only [interval_oc_eq_union, mem_union, mem_Ioc] +lemma not_mem_uIoc : a ∉ Ι b c ↔ a ≤ b ∧ a ≤ c ∨ c < a ∧ b < a := +by { simp only [uIoc_eq_union, mem_union, mem_Ioc, not_lt, ←not_le], tauto } -lemma not_mem_interval_oc : a ∉ Ι b c ↔ a ≤ b ∧ a ≤ c ∨ c < a ∧ b < a := -by { simp only [interval_oc_eq_union, mem_union, mem_Ioc, not_lt, ←not_le], tauto } +@[simp] lemma left_mem_uIoc : a ∈ Ι a b ↔ b < a := by simp [mem_uIoc] +@[simp] lemma right_mem_uIoc : b ∈ Ι a b ↔ a < b := by simp [mem_uIoc] -@[simp] lemma left_mem_interval_oc : a ∈ Ι a b ↔ b < a := by simp [mem_interval_oc] -@[simp] lemma right_mem_interval_oc : b ∈ Ι a b ↔ a < b := by simp [mem_interval_oc] - -lemma forall_interval_oc_iff {P : α → Prop} : +lemma forall_uIoc_iff {P : α → Prop} : (∀ x ∈ Ι a b, P x) ↔ (∀ x ∈ Ioc a b, P x) ∧ (∀ x ∈ Ioc b a, P x) := -by simp only [interval_oc_eq_union, mem_union, or_imp_distrib, forall_and_distrib] - -lemma interval_oc_subset_interval_oc_of_interval_subset_interval {a b c d : α} - (h : [a, b] ⊆ [c, d]) : Ι a b ⊆ Ι c d := -Ioc_subset_Ioc (interval_subset_interval_iff_le.1 h).1 (interval_subset_interval_iff_le.1 h).2 +by simp only [uIoc_eq_union, mem_union, or_imp_distrib, forall_and_distrib] -lemma interval_oc_swap (a b : α) : Ι a b = Ι b a := -by simp only [interval_oc, min_comm a b, max_comm a b] +lemma uIoc_subset_uIoc_of_uIcc_subset_uIcc (h : [a, b] ⊆ [c, d]) : Ι a b ⊆ Ι c d := +Ioc_subset_Ioc (uIcc_subset_uIcc_iff_le.1 h).1 (uIcc_subset_uIcc_iff_le.1 h).2 -lemma Ioc_subset_interval_oc : Ioc a b ⊆ Ι a b := -Ioc_subset_Ioc (min_le_left _ _) (le_max_right _ _) +lemma uIoc_swap (a b : α) : Ι a b = Ι b a := by simp only [uIoc, min_comm a b, max_comm a b] -lemma Ioc_subset_interval_oc' : Ioc a b ⊆ Ι b a := -Ioc_subset_Ioc (min_le_right _ _) (le_max_left _ _) +lemma Ioc_subset_uIoc : Ioc a b ⊆ Ι a b := Ioc_subset_Ioc (min_le_left _ _) (le_max_right _ _) +lemma Ioc_subset_uIoc' : Ioc a b ⊆ Ι b a := Ioc_subset_Ioc (min_le_right _ _) (le_max_left _ _) -lemma eq_of_mem_interval_oc_of_mem_interval_oc : a ∈ Ι b c → b ∈ Ι a c → a = b := -by simp_rw mem_interval_oc; rintro (⟨_, _⟩ | ⟨_, _⟩) (⟨_, _⟩ | ⟨_, _⟩); apply le_antisymm; +lemma eq_of_mem_uIoc_of_mem_uIoc : a ∈ Ι b c → b ∈ Ι a c → a = b := +by simp_rw mem_uIoc; rintro (⟨_, _⟩ | ⟨_, _⟩) (⟨_, _⟩ | ⟨_, _⟩); apply le_antisymm; assumption <|> exact le_of_lt ‹_› <|> exact le_trans ‹_› (le_of_lt ‹_›) -lemma eq_of_mem_interval_oc_of_mem_interval_oc' : b ∈ Ι a c → c ∈ Ι a b → b = c := -by simpa only [interval_oc_swap a] using eq_of_mem_interval_oc_of_mem_interval_oc +lemma eq_of_mem_uIoc_of_mem_uIoc' : b ∈ Ι a c → c ∈ Ι a b → b = c := +by simpa only [uIoc_swap a] using eq_of_mem_uIoc_of_mem_uIoc -lemma eq_of_not_mem_interval_oc_of_not_mem_interval_oc (ha : a ≤ c) (hb : b ≤ c) : +lemma eq_of_not_mem_uIoc_of_not_mem_uIoc (ha : a ≤ c) (hb : b ≤ c) : a ∉ Ι b c → b ∉ Ι a c → a = b := -by simp_rw not_mem_interval_oc; rintro (⟨_, _⟩ | ⟨_, _⟩) (⟨_, _⟩ | ⟨_, _⟩); apply le_antisymm; +by simp_rw not_mem_uIoc; rintro (⟨_, _⟩ | ⟨_, _⟩) (⟨_, _⟩ | ⟨_, _⟩); apply le_antisymm; assumption <|> exact le_of_lt ‹_› <|> cases not_le_of_lt ‹_› ‹_› -lemma interval_oc_injective_right (a : α) : injective (λ b, Ι b a) := +lemma uIoc_injective_right (a : α) : injective (λ b, Ι b a) := begin rintro b c h, rw ext_iff at h, obtain ha | ha := le_or_lt b a, { have hb := (h b).not, - simp only [ha, left_mem_interval_oc, not_lt, true_iff, not_mem_interval_oc, ←not_le, and_true, + simp only [ha, left_mem_uIoc, not_lt, true_iff, not_mem_uIoc, ←not_le, and_true, not_true, false_and, not_false_iff, true_iff, or_false] at hb, refine hb.eq_of_not_lt (λ hc, _), simpa [ha, and_iff_right hc, ←@not_le _ _ _ a, -not_le] using h c }, - { refine eq_of_mem_interval_oc_of_mem_interval_oc ((h _).1 $ left_mem_interval_oc.2 ha) - ((h _).2 $ left_mem_interval_oc.2 $ ha.trans_le _), - simpa [ha, ha.not_le, mem_interval_oc] using h b } + { refine eq_of_mem_uIoc_of_mem_uIoc ((h _).1 $ left_mem_uIoc.2 ha) + ((h _).2 $ left_mem_uIoc.2 $ ha.trans_le _), + simpa [ha, ha.not_le, mem_uIoc] using h b } end -lemma interval_oc_injective_left (a : α) : injective (Ι a) := -by simpa only [interval_oc_swap] using interval_oc_injective_right a +lemma uIoc_injective_left (a : α) : injective (Ι a) := +by simpa only [uIoc_swap] using uIoc_injective_right a end linear_order end set diff --git a/src/data/set/pointwise/interval.lean b/src/data/set/pointwise/interval.lean index b8f08d0378690..d62629b46dcdd 100644 --- a/src/data/set/pointwise/interval.lean +++ b/src/data/set/pointwise/interval.lean @@ -254,56 +254,56 @@ end ordered_add_comm_group section linear_ordered_add_comm_group variables [linear_ordered_add_comm_group α] (a b c d : α) -@[simp] lemma preimage_const_add_interval : (λ x, a + x) ⁻¹' [b, c] = [b - a, c - a] := +@[simp] lemma preimage_const_add_uIcc : (λ x, a + x) ⁻¹' [b, c] = [b - a, c - a] := by simp only [←Icc_min_max, preimage_const_add_Icc, min_sub_sub_right, max_sub_sub_right] -@[simp] lemma preimage_add_const_interval : (λ x, x + a) ⁻¹' [b, c] = [b - a, c - a] := -by simpa only [add_comm] using preimage_const_add_interval a b c +@[simp] lemma preimage_add_const_uIcc : (λ x, x + a) ⁻¹' [b, c] = [b - a, c - a] := +by simpa only [add_comm] using preimage_const_add_uIcc a b c -@[simp] lemma preimage_neg_interval : - [a, b] = [-a, -b] := +@[simp] lemma preimage_neg_uIcc : - [a, b] = [-a, -b] := by simp only [←Icc_min_max, preimage_neg_Icc, min_neg_neg, max_neg_neg] -@[simp] lemma preimage_sub_const_interval : (λ x, x - a) ⁻¹' [b, c] = [b + a, c + a] := +@[simp] lemma preimage_sub_const_uIcc : (λ x, x - a) ⁻¹' [b, c] = [b + a, c + a] := by simp [sub_eq_add_neg] -@[simp] lemma preimage_const_sub_interval : (λ x, a - x) ⁻¹' [b, c] = [a - b, a - c] := +@[simp] lemma preimage_const_sub_uIcc : (λ x, a - x) ⁻¹' [b, c] = [a - b, a - c] := by { simp_rw [←Icc_min_max, preimage_const_sub_Icc], simp only [sub_eq_add_neg, min_add_add_left, max_add_add_left, min_neg_neg, max_neg_neg], } -@[simp] lemma image_const_add_interval : (λ x, a + x) '' [b, c] = [a + b, a + c] := +@[simp] lemma image_const_add_uIcc : (λ x, a + x) '' [b, c] = [a + b, a + c] := by simp [add_comm] -@[simp] lemma image_add_const_interval : (λ x, x + a) '' [b, c] = [b + a, c + a] := +@[simp] lemma image_add_const_uIcc : (λ x, x + a) '' [b, c] = [b + a, c + a] := by simp -@[simp] lemma image_const_sub_interval : (λ x, a - x) '' [b, c] = [a - b, a - c] := +@[simp] lemma image_const_sub_uIcc : (λ x, a - x) '' [b, c] = [a - b, a - c] := by simp [sub_eq_add_neg, image_comp (λ x, a + x) (λ x, -x)] -@[simp] lemma image_sub_const_interval : (λ x, x - a) '' [b, c] = [b - a, c - a] := +@[simp] lemma image_sub_const_uIcc : (λ x, x - a) '' [b, c] = [b - a, c - a] := by simp [sub_eq_add_neg, add_comm] -lemma image_neg_interval : has_neg.neg '' [a, b] = [-a, -b] := by simp +lemma image_neg_uIcc : has_neg.neg '' [a, b] = [-a, -b] := by simp variables {a b c d} /-- If `[c, d]` is a subinterval of `[a, b]`, then the distance between `c` and `d` is less than or equal to that of `a` and `b` -/ -lemma abs_sub_le_of_subinterval (h : [c, d] ⊆ [a, b]) : |d - c| ≤ |b - a| := +lemma abs_sub_le_of_uIcc_subset_uIcc (h : [c, d] ⊆ [a, b]) : |d - c| ≤ |b - a| := begin rw [← max_sub_min_eq_abs, ← max_sub_min_eq_abs], - rw [interval_subset_interval_iff_le] at h, + rw [uIcc_subset_uIcc_iff_le] at h, exact sub_le_sub h.2 h.1, end /-- If `c ∈ [a, b]`, then the distance between `a` and `c` is less than or equal to that of `a` and `b` -/ -lemma abs_sub_left_of_mem_interval (h : c ∈ [a, b]) : |c - a| ≤ |b - a| := -abs_sub_le_of_subinterval (interval_subset_interval_left h) +lemma abs_sub_left_of_mem_uIcc (h : c ∈ [a, b]) : |c - a| ≤ |b - a| := +abs_sub_le_of_uIcc_subset_uIcc $ uIcc_subset_uIcc_left h /-- If `x ∈ [a, b]`, then the distance between `c` and `b` is less than or equal to that of `a` and `b` -/ -lemma abs_sub_right_of_mem_interval (h : c ∈ [a, b]) : |b - c| ≤ |b - a| := -abs_sub_le_of_subinterval (interval_subset_interval_right h) +lemma abs_sub_right_of_mem_uIcc (h : c ∈ [a, b]) : |b - c| ≤ |b - a| := +abs_sub_le_of_uIcc_subset_uIcc $ uIcc_subset_uIcc_right h end linear_ordered_add_comm_group @@ -442,32 +442,32 @@ by simpa only [mul_comm] using preimage_mul_const_Ico_of_neg a b h ((*) c) ⁻¹' (Icc a b) = Icc (b / c) (a / c) := by simpa only [mul_comm] using preimage_mul_const_Icc_of_neg a b h -@[simp] lemma preimage_mul_const_interval (ha : a ≠ 0) (b c : α) : +@[simp] lemma preimage_mul_const_uIcc (ha : a ≠ 0) (b c : α) : (λ x, x * a) ⁻¹' [b, c] = [b / a, c / a] := (lt_or_gt_of_ne ha).elim (λ h, by simp [←Icc_min_max, h, h.le, min_div_div_right_of_nonpos, max_div_div_right_of_nonpos]) (λ (ha : 0 < a), by simp [←Icc_min_max, ha, ha.le, min_div_div_right, max_div_div_right]) -@[simp] lemma preimage_const_mul_interval (ha : a ≠ 0) (b c : α) : +@[simp] lemma preimage_const_mul_uIcc (ha : a ≠ 0) (b c : α) : (λ x, a * x) ⁻¹' [b, c] = [b / a, c / a] := -by simp only [← preimage_mul_const_interval ha, mul_comm] +by simp only [← preimage_mul_const_uIcc ha, mul_comm] -@[simp] lemma preimage_div_const_interval (ha : a ≠ 0) (b c : α) : +@[simp] lemma preimage_div_const_uIcc (ha : a ≠ 0) (b c : α) : (λ x, x / a) ⁻¹' [b, c] = [b * a, c * a] := -by simp only [div_eq_mul_inv, preimage_mul_const_interval (inv_ne_zero ha), inv_inv] +by simp only [div_eq_mul_inv, preimage_mul_const_uIcc (inv_ne_zero ha), inv_inv] -@[simp] lemma image_mul_const_interval (a b c : α) : (λ x, x * a) '' [b, c] = [b * a, c * a] := +@[simp] lemma image_mul_const_uIcc (a b c : α) : (λ x, x * a) '' [b, c] = [b * a, c * a] := if ha : a = 0 then by simp [ha] else calc (λ x, x * a) '' [b, c] = (λ x, x * a⁻¹) ⁻¹' [b, c] : (units.mk0 a ha).mul_right.image_eq_preimage _ ... = (λ x, x / a) ⁻¹' [b, c] : by simp only [div_eq_mul_inv] -... = [b * a, c * a] : preimage_div_const_interval ha _ _ +... = [b * a, c * a] : preimage_div_const_uIcc ha _ _ -@[simp] lemma image_const_mul_interval (a b c : α) : (λ x, a * x) '' [b, c] = [a * b, a * c] := -by simpa only [mul_comm] using image_mul_const_interval a b c +@[simp] lemma image_const_mul_uIcc (a b c : α) : (λ x, a * x) '' [b, c] = [a * b, a * c] := +by simpa only [mul_comm] using image_mul_const_uIcc a b c -@[simp] lemma image_div_const_interval (a b c : α) : (λ x, x / a) '' [b, c] = [b / a, c / a] := -by simp only [div_eq_mul_inv, image_mul_const_interval] +@[simp] lemma image_div_const_uIcc (a b c : α) : (λ x, x / a) '' [b, c] = [b / a, c / a] := +by simp only [div_eq_mul_inv, image_mul_const_uIcc] lemma image_mul_right_Icc' (a b : α) {c : α} (h : 0 < c) : (λ x, x * c) '' Icc a b = Icc (a * c) (b * c) := diff --git a/src/linear_algebra/affine_space/affine_map.lean b/src/linear_algebra/affine_space/affine_map.lean index f6817af397589..be0d1ad6dd466 100644 --- a/src/linear_algebra/affine_space/affine_map.lean +++ b/src/linear_algebra/affine_space/affine_map.lean @@ -501,9 +501,9 @@ by rw decomp ; simp only [linear_map.map_zero, pi.add_apply, add_sub_cancel, zer omit V1 -lemma image_interval {k : Type*} [linear_ordered_field k] (f : k →ᵃ[k] k) +lemma image_uIcc {k : Type*} [linear_ordered_field k] (f : k →ᵃ[k] k) (a b : k) : - f '' set.interval a b = set.interval (f a) (f b) := + f '' set.uIcc a b = set.uIcc (f a) (f b) := begin have : ⇑f = (λ x, x + f 0) ∘ λ x, x * (f 1 - f 0), { ext x, @@ -511,7 +511,7 @@ begin rw [← f.linear_map_vsub, ← f.linear.map_smul, ← f.map_vadd], simp only [vsub_eq_sub, add_zero, mul_one, vadd_eq_add, sub_zero, smul_eq_mul] }, rw [this, set.image_comp], - simp only [set.image_add_const_interval, set.image_mul_const_interval] + simp only [set.image_add_const_uIcc, set.image_mul_const_uIcc] end section diff --git a/src/measure_theory/constructions/borel_space.lean b/src/measure_theory/constructions/borel_space.lean index 04f2a970078c0..e97b2770a683a 100644 --- a/src/measure_theory/constructions/borel_space.lean +++ b/src/measure_theory/constructions/borel_space.lean @@ -467,7 +467,7 @@ instance nhds_within_Iio_is_measurably_generated : (𝓝[Iio b] a).is_measurably_generated := measurable_set_Iio.nhds_within_is_measurably_generated _ -instance nhds_within_interval_is_measurably_generated : +instance nhds_within_uIcc_is_measurably_generated : is_measurably_generated (𝓝[[a, b]] x) := nhds_within_Icc_is_measurably_generated @@ -699,15 +699,10 @@ end linear_order section linear_order -variables [linear_order α] [order_closed_topology α] +variables [linear_order α] [order_closed_topology α] {a b : α} -@[measurability] -lemma measurable_set_interval {a b : α} : measurable_set (interval a b) := -measurable_set_Icc - -@[measurability] -lemma measurable_set_interval_oc {a b : α} : measurable_set (interval_oc a b) := -measurable_set_Ioc +@[measurability] lemma measurable_set_uIcc : measurable_set (uIcc a b) := measurable_set_Icc +@[measurability] lemma measurable_set_uIoc : measurable_set (uIoc a b) := measurable_set_Ioc variables [second_countable_topology α] diff --git a/src/measure_theory/function/locally_integrable.lean b/src/measure_theory/function/locally_integrable.lean index 1e5d2f178bcb2..fc98e48967e75 100644 --- a/src/measure_theory/function/locally_integrable.lean +++ b/src/measure_theory/function/locally_integrable.lean @@ -216,15 +216,15 @@ lemma continuous.integrable_on_Ioc [preorder X] [compact_Icc_space X] (hf : cont integrable_on f (Ioc a b) μ := hf.integrable_on_Icc.mono_set Ioc_subset_Icc_self -lemma continuous_on.integrable_on_interval [linear_order X] [compact_Icc_space X] +lemma continuous_on.integrable_on_uIcc [linear_order X] [compact_Icc_space X] (hf : continuous_on f [a, b]) : integrable_on f [a, b] μ := hf.integrable_on_Icc -lemma continuous.integrable_on_interval [linear_order X] [compact_Icc_space X] (hf : continuous f) : +lemma continuous.integrable_on_uIcc [linear_order X] [compact_Icc_space X] (hf : continuous f) : integrable_on f [a, b] μ := hf.integrable_on_Icc -lemma continuous.integrable_on_interval_oc [linear_order X] [compact_Icc_space X] +lemma continuous.integrable_on_uIoc [linear_order X] [compact_Icc_space X] (hf : continuous f) : integrable_on f (Ι a b) μ := hf.integrable_on_Ioc diff --git a/src/measure_theory/function/strongly_measurable/basic.lean b/src/measure_theory/function/strongly_measurable/basic.lean index efb7975a5578b..be7925460e9af 100644 --- a/src/measure_theory/function/strongly_measurable/basic.lean +++ b/src/measure_theory/function/strongly_measurable/basic.lean @@ -1603,12 +1603,12 @@ protected lemma Union [pseudo_metrizable_space β] {s : ι → set α} ae_strongly_measurable f (μ.restrict s) ∧ ae_strongly_measurable f (μ.restrict t) := by simp only [union_eq_Union, ae_strongly_measurable_Union_iff, bool.forall_bool, cond, and.comm] -lemma ae_strongly_measurable_interval_oc_iff [linear_order α] [pseudo_metrizable_space β] +lemma ae_strongly_measurable_uIoc_iff [linear_order α] [pseudo_metrizable_space β] {f : α → β} {a b : α} : - ae_strongly_measurable f (μ.restrict $ interval_oc a b) ↔ + ae_strongly_measurable f (μ.restrict $ uIoc a b) ↔ ae_strongly_measurable f (μ.restrict $ Ioc a b) ∧ ae_strongly_measurable f (μ.restrict $ Ioc b a) := -by rw [interval_oc_eq_union, ae_strongly_measurable_union_iff] +by rw [uIoc_eq_union, ae_strongly_measurable_union_iff] lemma smul_measure {R : Type*} [monoid R] [distrib_mul_action R ℝ≥0∞] [is_scalar_tower R ℝ≥0∞ ℝ≥0∞] (h : ae_strongly_measurable f μ) (c : R) : diff --git a/src/measure_theory/integral/circle_integral.lean b/src/measure_theory/integral/circle_integral.lean index 9cd190c8bdbcd..72912cc375d0b 100644 --- a/src/measure_theory/integral/circle_integral.lean +++ b/src/measure_theory/integral/circle_integral.lean @@ -259,7 +259,7 @@ begin { intro h, contrapose! h, rcases h with ⟨hR, hn, hw⟩, simp only [circle_integrable_iff R, deriv_circle_map], rw ← image_circle_map_Ioc at hw, rcases hw with ⟨θ, hθ, rfl⟩, - replace hθ : θ ∈ [0, 2 * π], from Icc_subset_interval (Ioc_subset_Icc_self hθ), + replace hθ : θ ∈ [0, 2 * π], from Icc_subset_uIcc (Ioc_subset_Icc_self hθ), refine not_interval_integrable_of_sub_inv_is_O_punctured _ real.two_pi_pos.ne hθ, set f : ℝ → ℂ := λ θ', circle_map c R θ' - circle_map c R θ, have : ∀ᶠ θ' in 𝓝[≠] θ, f θ' ∈ ball (0 : ℂ) 1 \ {0}, diff --git a/src/measure_theory/integral/circle_transform.lean b/src/measure_theory/integral/circle_transform.lean index 7ca20fb8487a1..1356aa0a5464c 100644 --- a/src/measure_theory/integral/circle_transform.lean +++ b/src/measure_theory/integral/circle_transform.lean @@ -128,9 +128,9 @@ lemma abs_circle_transform_bounding_function_le {R r : ℝ} (hr : r < R) (hr' : begin have cts := continuous_on_abs_circle_transform_bounding_function hr z, have comp : is_compact (closed_ball z r ×ˢ [0, 2 * π]), - { apply_rules [is_compact.prod, proper_space.is_compact_closed_ball z r, is_compact_interval], }, + { apply_rules [is_compact.prod, proper_space.is_compact_closed_ball z r, is_compact_uIcc], }, have none : (closed_ball z r ×ˢ [0, 2 * π]).nonempty := - (nonempty_closed_ball.2 hr').prod nonempty_interval, + (nonempty_closed_ball.2 hr').prod nonempty_uIcc, have := is_compact.exists_forall_ge comp none (cts.mono (by { intro z, simp only [mem_prod, mem_closed_ball, mem_univ, and_true, and_imp], tauto })), simpa only [set_coe.forall, subtype.coe_mk, set_coe.exists], @@ -158,7 +158,7 @@ begin obtain ⟨y1, hy1, hfun⟩ := periodic.exists_mem_Ico₀ (circle_transform_deriv_periodic R z v f) real.two_pi_pos y, have hy2: y1 ∈ [0, 2*π], by {convert (Ico_subset_Icc_self hy1), - simp [interval_of_le real.two_pi_pos.le]}, + simp [uIcc_of_le real.two_pi_pos.le]}, have := mul_le_mul (hab ⟨⟨v, y1⟩, ⟨ball_subset_closed_ball (H hv), hy2⟩⟩) (HX2 (circle_map z R y1) (circle_map_mem_sphere z hR.le y1)) (complex.abs.nonneg _) (complex.abs.nonneg _), diff --git a/src/measure_theory/integral/divergence_theorem.lean b/src/measure_theory/integral/divergence_theorem.lean index 8219c2b6bc01b..f458e7b18d655 100644 --- a/src/measure_theory/integral/divergence_theorem.lean +++ b/src/measure_theory/integral/divergence_theorem.lean @@ -404,9 +404,9 @@ theorem integral_eq_of_has_deriv_within_at_off_countable (f f' : ℝ → E) {a b ∫ x in a..b, f' x = f b - f a := begin cases le_total a b with hab hab, - { simp only [interval_of_le hab, min_eq_left hab, max_eq_right hab] at *, + { simp only [uIcc_of_le hab, min_eq_left hab, max_eq_right hab] at *, exact integral_eq_of_has_deriv_within_at_off_countable_of_le f f' hab hs Hc Hd Hi }, - { simp only [interval_of_ge hab, min_eq_right hab, max_eq_left hab] at *, + { simp only [uIcc_of_ge hab, min_eq_right hab, max_eq_left hab] at *, rw [interval_integral.integral_symm, neg_eq_iff_neg_eq, neg_sub, eq_comm], exact integral_eq_of_has_deriv_within_at_off_countable_of_le f f' hab hs Hc Hd Hi.symm } end @@ -488,7 +488,7 @@ lemma integral2_divergence_prod_of_has_fderiv_within_at_off_countable (f g : ℝ begin wlog h₁ : a₁ ≤ b₁ := le_total a₁ b₁ using [a₁ b₁, b₁ a₁] tactic.skip, wlog h₂ : a₂ ≤ b₂ := le_total a₂ b₂ using [a₂ b₂, b₂ a₂] tactic.skip, - { simp only [interval_of_le h₁, interval_of_le h₂, min_eq_left, max_eq_right, h₁, h₂] + { simp only [uIcc_of_le h₁, uIcc_of_le h₂, min_eq_left, max_eq_right, h₁, h₂] at Hcf Hcg Hdf Hdg Hi, calc ∫ x in a₁..b₁, ∫ y in a₂..b₂, f' (x, y) (1, 0) + g' (x, y) (0, 1) = ∫ x in Icc a₁ b₁, ∫ y in Icc a₂ b₂, f' (x, y) (1, 0) + g' (x, y) (0, 1) : @@ -503,11 +503,11 @@ begin apply integral_divergence_prod_Icc_of_has_fderiv_within_at_off_countable_of_le f g f' g' (a₁, a₂) (b₁, b₂) ⟨h₁, h₂⟩ s; assumption end }, - { rw [interval_swap b₂ a₂, min_comm b₂ a₂, max_comm b₂ a₂] at this, + { rw [uIcc_comm b₂ a₂, min_comm b₂ a₂, max_comm b₂ a₂] at this, intros Hcf Hcg Hdf Hdg Hi, simp only [interval_integral.integral_symm b₂ a₂, interval_integral.integral_neg], refine (congr_arg has_neg.neg (this Hcf Hcg Hdf Hdg Hi)).trans _, abel }, - { rw [interval_swap b₁ a₁, min_comm b₁ a₁, max_comm b₁ a₁] at this, + { rw [uIcc_comm b₁ a₁, min_comm b₁ a₁, max_comm b₁ a₁] at this, intros Hcf Hcg Hdf Hdg Hi, simp only [interval_integral.integral_symm b₁ a₁], refine (congr_arg has_neg.neg (this Hcf Hcg Hdf Hdg Hi)).trans _, abel } diff --git a/src/measure_theory/integral/integral_eq_improper.lean b/src/measure_theory/integral/integral_eq_improper.lean index c0814a8b2b830..43b38176df3eb 100644 --- a/src/measure_theory/integral/integral_eq_improper.lean +++ b/src/measure_theory/integral/integral_eq_improper.lean @@ -684,12 +684,12 @@ begin have i1 : Ioo (min a b) (max a b) ⊆ Ioi a, { rw min_eq_left hb.le, exact Ioo_subset_Ioi_self }, have i2 : [a, b] ⊆ Ici a, - { rw interval_of_le hb.le, exact Icc_subset_Ici_self }, + { rw uIcc_of_le hb.le, exact Icc_subset_Ici_self }, refine interval_integral.integral_comp_smul_deriv''' (hf.mono i2) (λ x hx, hff' x $ mem_of_mem_of_subset hx i1) (hg_cont.mono $ image_subset _ _) (hg1.mono_set $ image_subset _ _) (hg2.mono_set i2), { rw min_eq_left hb.le, exact Ioo_subset_Ioi_self }, - { rw interval_of_le hb.le, exact Icc_subset_Ici_self } }, + { rw uIcc_of_le hb.le, exact Icc_subset_Ici_self } }, rw integrable_on_Ici_iff_integrable_on_Ioi at hg2, have t2 := interval_integral_tendsto_integral_Ioi _ hg2 tendsto_id, have : Ioi (f a) ⊆ f '' Ici a := (Ioi_subset_Ici_self.trans $ diff --git a/src/measure_theory/integral/interval_average.lean b/src/measure_theory/integral/interval_average.lean index f1d2f75d36789..7b7f651c4ddd7 100644 --- a/src/measure_theory/integral/interval_average.lean +++ b/src/measure_theory/integral/interval_average.lean @@ -33,14 +33,14 @@ notation `⨍` binders ` in ` a `..` b `, ` r:(scoped:60 f, average (measure.restrict volume (Ι a b)) f) := r lemma interval_average_symm (f : ℝ → E) (a b : ℝ) : ⨍ x in a..b, f x = ⨍ x in b..a, f x := -by rw [set_average_eq, set_average_eq, interval_oc_swap] +by rw [set_average_eq, set_average_eq, uIoc_swap] lemma interval_average_eq (f : ℝ → E) (a b : ℝ) : ⨍ x in a..b, f x = (b - a)⁻¹ • ∫ x in a..b, f x := begin cases le_or_lt a b with h h, - { rw [set_average_eq, interval_oc_of_le h, real.volume_Ioc, interval_integral.integral_of_le h, + { rw [set_average_eq, uIoc_of_le h, real.volume_Ioc, interval_integral.integral_of_le h, ennreal.to_real_of_real (sub_nonneg.2 h)] }, - { rw [set_average_eq, interval_oc_of_lt h, real.volume_Ioc, interval_integral.integral_of_ge h.le, + { rw [set_average_eq, uIoc_of_lt h, real.volume_Ioc, interval_integral.integral_of_ge h.le, ennreal.to_real_of_real (sub_nonneg.2 h.le), smul_neg, ← neg_smul, ← inv_neg, neg_sub] } end diff --git a/src/measure_theory/integral/interval_integral.lean b/src/measure_theory/integral/interval_integral.lean index 20f10a469b7fd..925d890f17376 100644 --- a/src/measure_theory/integral/interval_integral.lean +++ b/src/measure_theory/integral/interval_integral.lean @@ -106,7 +106,7 @@ a context with the stronger assumption that `f'` is continuous, one can use In order to avoid `if`s in the definition, we define `interval_integrable f μ a b` as `integrable_on f (Ioc a b) μ ∧ integrable_on f (Ioc b a) μ`. For any `a`, `b` one of these -intervals is empty and the other coincides with `set.interval_oc a b = set.Ioc (min a b) (max a b)`. +intervals is empty and the other coincides with `set.uIoc a b = set.Ioc (min a b) (max a b)`. Similarly, we define `∫ x in a..b, f x ∂μ` to be `∫ x in Ioc a b, f x ∂μ - ∫ x in Ioc b a, f x ∂μ`. Again, for any `a`, `b` one of these integrals is zero, and the other gives the expected result. @@ -116,7 +116,7 @@ the cases `a ≤ b` and `b ≤ a` separately. ### Choice of the interval -We use integral over `set.interval_oc a b = set.Ioc (min a b) (max a b)` instead of one of the other +We use integral over `set.uIoc a b = set.Ioc (min a b) (max a b)` instead of one of the other three possible intervals with the same endpoints for two reasons: * this way `∫ x in a..b, f x ∂μ + ∫ x in b..c, f x ∂μ = ∫ x in a..c, f x ∂μ` holds whenever @@ -151,7 +151,7 @@ atom at one of the endpoints. There are some `FTC_filter` instances where the fact that it is one-sided or two-sided depends on the point, namely `(x, 𝓝[Icc a b] x, 𝓝[Icc a b] x)` -(resp. `(x, 𝓝[[a, b]] x, 𝓝[[a, b]] x)`, where `[a, b] = set.interval a b`), +(resp. `(x, 𝓝[[a, b]] x, 𝓝[[a, b]] x)`, where `[a, b] = set.uIcc a b`), with `x ∈ Icc a b` (resp. `x ∈ [a, b]`). This results in a two-sided derivatives for `x ∈ Ioo a b` and one-sided derivatives for `x ∈ {a, b}`. Other instances could be added when needed (in that case, one also needs to add @@ -186,19 +186,19 @@ section variables {f : ℝ → E} {a b : ℝ} {μ : measure ℝ} /-- A function is interval integrable with respect to a given measure `μ` on `a..b` if and - only if it is integrable on `interval_oc a b` with respect to `μ`. This is an equivalent + only if it is integrable on `uIoc a b` with respect to `μ`. This is an equivalent definition of `interval_integrable`. -/ lemma interval_integrable_iff : interval_integrable f μ a b ↔ integrable_on f (Ι a b) μ := -by rw [interval_oc_eq_union, integrable_on_union, interval_integrable] +by rw [uIoc_eq_union, integrable_on_union, interval_integrable] /-- If a function is interval integrable with respect to a given measure `μ` on `a..b` then - it is integrable on `interval_oc a b` with respect to `μ`. -/ + it is integrable on `uIoc a b` with respect to `μ`. -/ lemma interval_integrable.def (h : interval_integrable f μ a b) : integrable_on f (Ι a b) μ := interval_integrable_iff.mp h lemma interval_integrable_iff_integrable_Ioc_of_le (hab : a ≤ b) : interval_integrable f μ a b ↔ integrable_on f (Ioc a b) μ := -by rw [interval_integrable_iff, interval_oc_of_le hab] +by rw [interval_integrable_iff, uIoc_of_le hab] lemma integrable_on_Icc_iff_integrable_on_Ioc' {f : ℝ → E} (ha : μ {a} ≠ ∞) : integrable_on f (Icc a b) μ ↔ integrable_on f (Ioc a b) μ := @@ -234,8 +234,8 @@ lemma integrable_on_Icc_iff_integrable_on_Ioo [has_no_atoms μ] {f : ℝ → E} by rw [integrable_on_Icc_iff_integrable_on_Ioc, integrable_on_Ioc_iff_integrable_on_Ioo] lemma interval_integrable_iff' [has_no_atoms μ] : - interval_integrable f μ a b ↔ integrable_on f (interval a b) μ := -by rw [interval_integrable_iff, ←Icc_min_max, interval_oc, integrable_on_Icc_iff_integrable_on_Ioc] + interval_integrable f μ a b ↔ integrable_on f (uIcc a b) μ := +by rw [interval_integrable_iff, ←Icc_min_max, uIoc, integrable_on_Icc_iff_integrable_on_Ioc] lemma interval_integrable_iff_integrable_Icc_of_le {f : ℝ → E} {a b : ℝ} (hab : a ≤ b) {μ : measure ℝ} [has_no_atoms μ] : @@ -255,15 +255,15 @@ lemma integrable_on_Ici_iff_integrable_on_Ioi [has_no_atoms μ] {f : ℝ → E} integrable_on_Ici_iff_integrable_on_Ioi' (by simp) /-- If a function is integrable with respect to a given measure `μ` then it is interval integrable - with respect to `μ` on `interval a b`. -/ + with respect to `μ` on `uIcc a b`. -/ lemma measure_theory.integrable.interval_integrable (hf : integrable f μ) : interval_integrable f μ a b := ⟨hf.integrable_on, hf.integrable_on⟩ lemma measure_theory.integrable_on.interval_integrable (hf : integrable_on f [a, b] μ) : interval_integrable f μ a b := -⟨measure_theory.integrable_on.mono_set hf (Ioc_subset_Icc_self.trans Icc_subset_interval), - measure_theory.integrable_on.mono_set hf (Ioc_subset_Icc_self.trans Icc_subset_interval')⟩ +⟨measure_theory.integrable_on.mono_set hf (Ioc_subset_Icc_self.trans Icc_subset_uIcc), + measure_theory.integrable_on.mono_set hf (Ioc_subset_Icc_self.trans Icc_subset_uIcc')⟩ lemma interval_integrable_const_iff {c : E} : interval_integrable (λ _, c) μ a b ↔ c = 0 ∨ μ (Ι a b) < ∞ := @@ -326,7 +326,7 @@ h.norm lemma mono (hf : interval_integrable f ν a b) (h1 : [c, d] ⊆ [a, b]) (h2 : μ ≤ ν) : interval_integrable f μ c d := interval_integrable_iff.mpr $ hf.def.mono - (interval_oc_subset_interval_oc_of_interval_subset_interval h1) h2 + (uIoc_subset_uIoc_of_uIcc_subset_uIcc h1) h2 lemma mono_measure (hf : interval_integrable f ν a b) (h : μ ≤ ν) : interval_integrable f μ a b := @@ -388,7 +388,7 @@ lemma mul_continuous_on {f g : ℝ → A} interval_integrable (λ x, f x * g x) μ a b := begin rw interval_integrable_iff at hf ⊢, - exact hf.mul_continuous_on_of_subset hg measurable_set_Ioc is_compact_interval Ioc_subset_Icc_self + exact hf.mul_continuous_on_of_subset hg measurable_set_Ioc is_compact_uIcc Ioc_subset_Icc_self end lemma continuous_on_mul {f g : ℝ → A} @@ -396,7 +396,7 @@ lemma continuous_on_mul {f g : ℝ → A} interval_integrable (λ x, g x * f x) μ a b := begin rw interval_integrable_iff at hf ⊢, - exact hf.continuous_on_mul_of_subset hg is_compact_interval measurable_set_Ioc Ioc_subset_Icc_self + exact hf.continuous_on_mul_of_subset hg is_compact_uIcc measurable_set_Ioc Ioc_subset_Icc_self end lemma const_mul {f : ℝ → A} @@ -419,7 +419,7 @@ begin ←integrable_on, measurable_embedding.integrable_on_map_iff A], convert hf using 1, { ext, simp only [comp_app], congr' 1, field_simp, ring }, - { rw preimage_mul_const_interval (inv_ne_zero hc), field_simp [hc] }, + { rw preimage_mul_const_uIcc (inv_ne_zero hc), field_simp [hc] }, end lemma iff_comp_neg : @@ -435,12 +435,12 @@ section variables {μ : measure ℝ} [is_locally_finite_measure μ] lemma continuous_on.interval_integrable {u : ℝ → E} {a b : ℝ} - (hu : continuous_on u (interval a b)) : interval_integrable u μ a b := + (hu : continuous_on u (uIcc a b)) : interval_integrable u μ a b := (continuous_on.integrable_on_Icc hu).interval_integrable lemma continuous_on.interval_integrable_of_Icc {u : ℝ → E} {a b : ℝ} (h : a ≤ b) (hu : continuous_on u (Icc a b)) : interval_integrable u μ a b := -continuous_on.interval_integrable ((interval_of_le h).symm ▸ hu) +continuous_on.interval_integrable ((uIcc_of_le h).symm ▸ hu) /-- A continuous function on `ℝ` is `interval_integrable` with respect to any locally finite measure `ν` on ℝ. -/ @@ -455,14 +455,14 @@ section variables {μ : measure ℝ} [is_locally_finite_measure μ] [conditionally_complete_linear_order E] [order_topology E] [second_countable_topology E] -lemma monotone_on.interval_integrable {u : ℝ → E} {a b : ℝ} (hu : monotone_on u (interval a b)) : +lemma monotone_on.interval_integrable {u : ℝ → E} {a b : ℝ} (hu : monotone_on u (uIcc a b)) : interval_integrable u μ a b := begin rw interval_integrable_iff, - exact (hu.integrable_on_is_compact is_compact_interval).mono_set Ioc_subset_Icc_self, + exact (hu.integrable_on_is_compact is_compact_uIcc).mono_set Ioc_subset_Icc_self, end -lemma antitone_on.interval_integrable {u : ℝ → E} {a b : ℝ} (hu : antitone_on u (interval a b)) : +lemma antitone_on.interval_integrable {u : ℝ → E} {a b : ℝ} (hu : antitone_on u (uIcc a b)) : interval_integrable u μ a b := hu.dual_right.interval_integrable @@ -550,18 +550,18 @@ by simp only [interval_integral, neg_sub] lemma integral_of_ge (h : b ≤ a) : ∫ x in a..b, f x ∂μ = -∫ x in Ioc b a, f x ∂μ := by simp only [integral_symm b, integral_of_le h] -lemma interval_integral_eq_integral_interval_oc (f : ℝ → E) (a b : ℝ) (μ : measure ℝ) : +lemma interval_integral_eq_integral_uIoc (f : ℝ → E) (a b : ℝ) (μ : measure ℝ) : ∫ x in a..b, f x ∂μ = (if a ≤ b then 1 else -1 : ℝ) • ∫ x in Ι a b, f x ∂μ := begin split_ifs with h, - { simp only [integral_of_le h, interval_oc_of_le h, one_smul] }, - { simp only [integral_of_ge (not_le.1 h).le, interval_oc_of_lt (not_le.1 h), neg_one_smul] } + { simp only [integral_of_le h, uIoc_of_le h, one_smul] }, + { simp only [integral_of_ge (not_le.1 h).le, uIoc_of_lt (not_le.1 h), neg_one_smul] } end lemma norm_interval_integral_eq (f : ℝ → E) (a b : ℝ) (μ : measure ℝ) : ‖∫ x in a..b, f x ∂μ‖ = ‖∫ x in Ι a b, f x ∂μ‖ := begin - simp_rw [interval_integral_eq_integral_interval_oc, norm_smul], + simp_rw [interval_integral_eq_integral_uIoc, norm_smul], split_ifs; simp only [norm_neg, norm_one, one_mul], end @@ -571,7 +571,7 @@ norm_interval_integral_eq f a b μ lemma integral_cases (f : ℝ → E) (a b) : ∫ x in a..b, f x ∂μ ∈ ({∫ x in Ι a b, f x ∂μ, -∫ x in Ι a b, f x ∂μ} : set E) := -by { rw interval_integral_eq_integral_interval_oc, split_ifs; simp } +by { rw interval_integral_eq_integral_uIoc, split_ifs; simp } lemma integral_undef (h : ¬ interval_integrable f μ a b) : ∫ x in a..b, f x ∂μ = 0 := @@ -588,12 +588,12 @@ by { contrapose! h, exact interval_integral.integral_undef h } lemma integral_non_ae_strongly_measurable (hf : ¬ ae_strongly_measurable f (μ.restrict (Ι a b))) : ∫ x in a..b, f x ∂μ = 0 := -by rw [interval_integral_eq_integral_interval_oc, integral_non_ae_strongly_measurable hf, smul_zero] +by rw [interval_integral_eq_integral_uIoc, integral_non_ae_strongly_measurable hf, smul_zero] lemma integral_non_ae_strongly_measurable_of_le (h : a ≤ b) (hf : ¬ ae_strongly_measurable f (μ.restrict (Ioc a b))) : ∫ x in a..b, f x ∂μ = 0 := -integral_non_ae_strongly_measurable $ by rwa [interval_oc_of_le h] +integral_non_ae_strongly_measurable $ by rwa [uIoc_of_le h] lemma norm_integral_min_max (f : ℝ → E) : ‖∫ x in min a b..max a b, f x ∂μ‖ = ‖∫ x in a..b, f x ∂μ‖ := @@ -601,9 +601,9 @@ by cases le_total a b; simp [*, integral_symm a b] lemma norm_integral_eq_norm_integral_Ioc (f : ℝ → E) : ‖∫ x in a..b, f x ∂μ‖ = ‖∫ x in Ι a b, f x ∂μ‖ := -by rw [← norm_integral_min_max, integral_of_le min_le_max, interval_oc] +by rw [← norm_integral_min_max, integral_of_le min_le_max, uIoc] -lemma abs_integral_eq_abs_integral_interval_oc (f : ℝ → ℝ) : +lemma abs_integral_eq_abs_integral_uIoc (f : ℝ → ℝ) : |∫ x in a..b, f x ∂μ| = |∫ x in Ι a b, f x ∂μ| := norm_integral_eq_norm_integral_Ioc f @@ -622,7 +622,7 @@ end lemma norm_integral_le_integral_norm (h : a ≤ b) : ‖∫ x in a..b, f x ∂μ‖ ≤ ∫ x in a..b, ‖f x‖ ∂μ := -norm_integral_le_integral_norm_Ioc.trans_eq $ by rw [interval_oc_of_le h, integral_of_le h] +norm_integral_le_integral_norm_Ioc.trans_eq $ by rw [uIoc_of_le h, integral_of_le h] lemma norm_integral_le_of_norm_le {g : ℝ → ℝ} (h : ∀ᵐ t ∂(μ.restrict $ Ι a b), ‖f t‖ ≤ g t) @@ -649,12 +649,12 @@ norm_integral_le_of_norm_le_const_ae $ eventually_of_forall h @[simp] lemma integral_add (hf : interval_integrable f μ a b) (hg : interval_integrable g μ a b) : ∫ x in a..b, f x + g x ∂μ = ∫ x in a..b, f x ∂μ + ∫ x in a..b, g x ∂μ := -by simp only [interval_integral_eq_integral_interval_oc, integral_add hf.def hg.def, smul_add] +by simp only [interval_integral_eq_integral_uIoc, integral_add hf.def hg.def, smul_add] lemma integral_finset_sum {ι} {s : finset ι} {f : ι → ℝ → E} (h : ∀ i ∈ s, interval_integrable (f i) μ a b) : ∫ x in a..b, ∑ i in s, f i x ∂μ = ∑ i in s, ∫ x in a..b, f i x ∂μ := -by simp only [interval_integral_eq_integral_interval_oc, +by simp only [interval_integral_eq_integral_uIoc, integral_finset_sum s (λ i hi, (h i hi).def), finset.smul_sum] @[simp] lemma integral_neg : ∫ x in a..b, -f x ∂μ = -∫ x in a..b, f x ∂μ := @@ -671,7 +671,7 @@ by simp only [interval_integral, integral_smul, smul_sub] @[simp] lemma integral_smul_const {𝕜 : Type*} [is_R_or_C 𝕜] [normed_space 𝕜 E] (f : ℝ → 𝕜) (c : E) : ∫ x in a..b, f x • c ∂μ = (∫ x in a..b, f x ∂μ) • c := -by simp only [interval_integral_eq_integral_interval_oc, integral_smul_const, smul_assoc] +by simp only [interval_integral_eq_integral_uIoc, integral_smul_const, smul_assoc] @[simp] lemma integral_const_mul {𝕜 : Type*} [is_R_or_C 𝕜] (r : 𝕜) (f : ℝ → 𝕜) : ∫ x in a..b, r * f x ∂μ = r * ∫ x in a..b, f x ∂μ := @@ -713,7 +713,7 @@ open continuous_linear_map lemma _root_.continuous_linear_map.interval_integral_apply {a b : ℝ} {φ : ℝ → F →L[𝕜] E} (hφ : interval_integrable φ μ a b) (v : F) : (∫ x in a..b, φ x ∂μ) v = ∫ x in a..b, φ x v ∂μ := -by simp_rw [interval_integral_eq_integral_interval_oc, ← integral_apply hφ.def v, coe_smul', +by simp_rw [interval_integral_eq_integral_uIoc, ← integral_apply hφ.def v, coe_smul', pi.smul_apply] variables [normed_space ℝ F] [complete_space F] @@ -978,7 +978,7 @@ by simp only [interval_integral, set_integral_congr_ae (measurable_set_Ioc) h, lemma integral_congr_ae (h : ∀ᵐ x ∂μ, x ∈ Ι a b → f x = g x) : ∫ x in a..b, f x ∂μ = ∫ x in a..b, g x ∂μ := -integral_congr_ae' (ae_interval_oc_iff.mp h).1 (ae_interval_oc_iff.mp h).2 +integral_congr_ae' (ae_uIoc_iff.mp h).1 (ae_uIoc_iff.mp h).2 lemma integral_zero_ae (h : ∀ᵐ x ∂μ, x ∈ Ι a b → f x = 0) : ∫ x in a..b, f x ∂μ = 0 := @@ -1004,8 +1004,8 @@ lemma tendsto_integral_filter_of_dominated_convergence {ι} {l : filter ι} (h_lim : ∀ᵐ x ∂μ, x ∈ Ι a b → tendsto (λ n, F n x) l (𝓝 (f x))) : tendsto (λn, ∫ x in a..b, F n x ∂μ) l (𝓝 $ ∫ x in a..b, f x ∂μ) := begin - simp only [interval_integrable_iff, interval_integral_eq_integral_interval_oc, - ← ae_restrict_iff' measurable_set_interval_oc] at *, + simp only [interval_integrable_iff, interval_integral_eq_integral_uIoc, + ← ae_restrict_iff' measurable_set_uIoc] at *, exact tendsto_const_nhds.smul (tendsto_integral_filter_of_dominated_convergence bound hF_meas h_bound bound_integrable h_lim) end @@ -1020,8 +1020,8 @@ lemma has_sum_integral_of_dominated_convergence {ι} [countable ι] (h_lim : ∀ᵐ t ∂μ, t ∈ Ι a b → has_sum (λ n, F n t) (f t)) : has_sum (λn, ∫ t in a..b, F n t ∂μ) (∫ t in a..b, f t ∂μ) := begin - simp only [interval_integrable_iff, interval_integral_eq_integral_interval_oc, - ← ae_restrict_iff' measurable_set_interval_oc] at *, + simp only [interval_integrable_iff, interval_integral_eq_integral_uIoc, + ← ae_restrict_iff' measurable_set_uIoc] at *, exact (has_sum_integral_of_dominated_convergence bound hF_meas h_bound bound_summable bound_integrable h_lim).const_smul end @@ -1091,7 +1091,7 @@ begin have h_int' : ∀ {x}, x ∈ Icc b₁ b₂ → interval_integrable f μ b₁ x, { rintros x ⟨h₁, h₂⟩, apply h_int.mono_set, - apply interval_subset_interval, + apply uIcc_subset_uIcc, { exact ⟨min_le_of_left_le (min_le_right a b₁), h₁.trans (h₂.trans $ le_max_of_le_right $ le_max_right _ _)⟩ }, { exact ⟨min_le_of_left_le $ (min_le_right _ _).trans h₁, @@ -1100,7 +1100,7 @@ begin { rintros b ⟨h₁, h₂⟩, rw ← integral_add_adjacent_intervals _ (h_int' ⟨h₁, h₂⟩), apply h_int.mono_set, - apply interval_subset_interval, + apply uIcc_subset_uIcc, { exact ⟨min_le_of_left_le (min_le_left a b₁), le_max_of_le_right (le_max_left _ _)⟩ }, { exact ⟨min_le_of_left_le (min_le_right _ _), le_max_of_le_right (h₁.trans $ h₂.trans (le_max_right a b₂))⟩ } }, @@ -1155,7 +1155,7 @@ begin by_cases h : a ≤ b, { have : ∀ x ∈ Icc a b, ∫ t in Ioc a x, f t ∂μ = ∫ t in a..x, f t ∂μ, { intros x x_in, - simp_rw [← interval_oc_of_le h, integral_of_le x_in.1] }, + simp_rw [← uIoc_of_le h, integral_of_le x_in.1] }, rw continuous_on_congr this, intros x₀ hx₀, refine continuous_within_at_primitive (measure_singleton x₀) _, @@ -1181,19 +1181,19 @@ begin intros b₀ hb₀, refine continuous_within_at_primitive (measure_singleton _) _, rw [min_eq_right ha.1, max_eq_right ha.2], - simpa [interval_integrable_iff, interval_oc] using h_int, + simpa [interval_integrable_iff, uIoc] using h_int, end lemma continuous_on_primitive_interval [has_no_atoms μ] - (h_int : integrable_on f (interval a b) μ) : - continuous_on (λ x, ∫ t in a..x, f t ∂ μ) (interval a b) := -continuous_on_primitive_interval' h_int.interval_integrable left_mem_interval + (h_int : integrable_on f (uIcc a b) μ) : + continuous_on (λ x, ∫ t in a..x, f t ∂ μ) (uIcc a b) := +continuous_on_primitive_interval' h_int.interval_integrable left_mem_uIcc lemma continuous_on_primitive_interval_left [has_no_atoms μ] - (h_int : integrable_on f (interval a b) μ) : - continuous_on (λ x, ∫ t in x..b, f t ∂ μ) (interval a b) := + (h_int : integrable_on f (uIcc a b) μ) : + continuous_on (λ x, ∫ t in x..b, f t ∂ μ) (uIcc a b) := begin - rw interval_swap a b at h_int ⊢, + rw uIcc_comm a b at h_int ⊢, simp only [integral_symm b], exact (continuous_on_primitive_interval h_int).neg, end @@ -1236,7 +1236,7 @@ begin { rw [integral_symm, neg_eq_zero, integral_eq_zero_iff_of_le_of_nonneg_ae hab hf hfi.symm] } end -/-- If `f` is nonnegative and integrable on the unordered interval `set.interval_oc a b`, then its +/-- If `f` is nonnegative and integrable on the unordered interval `set.uIoc a b`, then its integral over `a..b` is positive if and only if `a < b` and the measure of `function.support f ∩ set.Ioc a b` is positive. -/ lemma integral_pos_iff_support_of_nonneg_ae' @@ -1244,17 +1244,17 @@ lemma integral_pos_iff_support_of_nonneg_ae' 0 < ∫ x in a..b, f x ∂μ ↔ a < b ∧ 0 < μ (support f ∩ Ioc a b) := begin cases lt_or_le a b with hab hba, - { rw interval_oc_of_le hab.le at hf, + { rw uIoc_of_le hab.le at hf, simp only [hab, true_and, integral_of_le hab.le, set_integral_pos_iff_support_of_nonneg_ae hf hfi.1] }, { suffices : ∫ x in a..b, f x ∂μ ≤ 0, by simp only [this.not_lt, hba.not_lt, false_and], rw [integral_of_ge hba, neg_nonpos], - rw [interval_oc_swap, interval_oc_of_le hba] at hf, + rw [uIoc_swap, uIoc_of_le hba] at hf, exact integral_nonneg_of_ae hf } end /-- If `f` is nonnegative a.e.-everywhere and it is integrable on the unordered interval -`set.interval_oc a b`, then its integral over `a..b` is positive if and only if `a < b` and the +`set.uIoc a b`, then its integral over `a..b` is positive if and only if `a < b` and the measure of `function.support f ∩ set.Ioc a b` is positive. -/ lemma integral_pos_iff_support_of_nonneg_ae (hf : 0 ≤ᵐ[μ] f) (hfi : interval_integrable f μ a b) : 0 < ∫ x in a..b, f x ∂μ ↔ a < b ∧ 0 < μ (support f ∩ Ioc a b) := @@ -1365,11 +1365,11 @@ lemma abs_integral_mono_interval {c d } (h : Ι a b ⊆ Ι c d) (hf : 0 ≤ᵐ[μ.restrict (Ι c d)] f) (hfi : interval_integrable f μ c d) : |∫ x in a..b, f x ∂μ| ≤ |∫ x in c..d, f x ∂μ| := have hf' : 0 ≤ᵐ[μ.restrict (Ι a b)] f, from ae_mono (measure.restrict_mono h le_rfl) hf, -calc |∫ x in a..b, f x ∂μ| = |∫ x in Ι a b, f x ∂μ| : abs_integral_eq_abs_integral_interval_oc f +calc |∫ x in a..b, f x ∂μ| = |∫ x in Ι a b, f x ∂μ| : abs_integral_eq_abs_integral_uIoc f ... = ∫ x in Ι a b, f x ∂μ : abs_of_nonneg (measure_theory.integral_nonneg_of_ae hf') ... ≤ ∫ x in Ι c d, f x ∂μ : set_integral_mono_set hfi.def hf h.eventually_le ... ≤ |∫ x in Ι c d, f x ∂μ| : le_abs_self _ -... = |∫ x in c..d, f x ∂μ| : (abs_integral_eq_abs_integral_interval_oc f).symm +... = |∫ x in c..d, f x ∂μ| : (abs_integral_eq_abs_integral_uIoc f).symm end mono @@ -1456,7 +1456,7 @@ instance nhds_Icc {x a b : ℝ} [h : fact (x ∈ Icc a b)] : { pure_le := pure_le_nhds_within h.out, le_nhds := inf_le_left } -instance nhds_interval {x a b : ℝ} [h : fact (x ∈ [a, b])] : +instance nhds_uIcc {x a b : ℝ} [h : fact (x ∈ [a, b])] : FTC_filter x (𝓝[[a, b]] x) (𝓝[[a, b]] x) := by { haveI : fact (x ∈ set.Icc (min a b) (max a b)) := h, exact FTC_filter.nhds_Icc } @@ -2190,7 +2190,7 @@ begin -- the set `s` of points where this property holds is closed. have s_closed : is_closed s, { have : continuous_on (λ t, (g t - g a, ∫ u in a..t, (G' u).to_real)) (Icc a b), - { rw ← interval_of_le hab at G'int ⊢ hcont, + { rw ← uIcc_of_le hab at G'int ⊢ hcont, exact (hcont.sub continuous_on_const).prod (continuous_on_primitive_interval G'int) }, simp only [s, inter_comm], exact this.preimage_closed_of_closed is_closed_Icc order_closed_topology.is_closed_le' }, @@ -2292,7 +2292,7 @@ begin set s := {t | g b - g t ≤ ∫ u in t..b, φ u} ∩ Icc a b, have s_closed : is_closed s, { have : continuous_on (λ t, (g b - g t, ∫ u in t..b, φ u)) (Icc a b), - { rw ← interval_of_le hab at ⊢ hcont φint, + { rw ← uIcc_of_le hab at ⊢ hcont φint, exact (continuous_on_const.sub hcont).prod (continuous_on_primitive_interval_left φint) }, simp only [s, inter_comm], exact this.preimage_closed_of_closed is_closed_Icc is_closed_le_prod, }, @@ -2353,15 +2353,15 @@ end /-- Fundamental theorem of calculus-2: If `f : ℝ → E` is continuous on `[a, b]` and has a right derivative at `f' x` for all `x` in `[a, b)`, and `f'` is integrable on `[a, b]` then `∫ y in a..b, f' y` equals `f b - f a`. -/ -theorem integral_eq_sub_of_has_deriv_right (hcont : continuous_on f (interval a b)) +theorem integral_eq_sub_of_has_deriv_right (hcont : continuous_on f (uIcc a b)) (hderiv : ∀ x ∈ Ioo (min a b) (max a b), has_deriv_within_at f (f' x) (Ioi x) x) (hint : interval_integrable f' volume a b) : ∫ y in a..b, f' y = f b - f a := begin cases le_total a b with hab hab, - { simp only [interval_of_le, min_eq_left, max_eq_right, hab] at hcont hderiv hint, + { simp only [uIcc_of_le, min_eq_left, max_eq_right, hab] at hcont hderiv hint, apply integral_eq_sub_of_has_deriv_right_of_le hab hcont hderiv hint }, - { simp only [interval_of_ge, min_eq_right, max_eq_left, hab] at hcont hderiv, + { simp only [uIcc_of_ge, min_eq_right, max_eq_left, hab] at hcont hderiv, rw [integral_symm, integral_eq_sub_of_has_deriv_right_of_le hab hcont hderiv hint.symm, neg_sub] } end @@ -2378,7 +2378,7 @@ integral_eq_sub_of_has_deriv_right_of_le hab hcont (λ x hx, (hderiv x hx).has_d /-- Fundamental theorem of calculus-2: If `f : ℝ → E` has a derivative at `f' x` for all `x` in `[a, b]` and `f'` is integrable on `[a, b]`, then `∫ y in a..b, f' y` equals `f b - f a`. -/ theorem integral_eq_sub_of_has_deriv_at - (hderiv : ∀ x ∈ interval a b, has_deriv_at f (f' x) x) + (hderiv : ∀ x ∈ uIcc a b, has_deriv_at f (f' x) x) (hint : interval_integrable f' volume a b) : ∫ y in a..b, f' y = f b - f a := integral_eq_sub_of_has_deriv_right (has_deriv_at.continuous_on hderiv) @@ -2407,14 +2407,14 @@ end /-- Fundamental theorem of calculus-2: If `f : ℝ → E` is differentiable at every `x` in `[a, b]` and its derivative is integrable on `[a, b]`, then `∫ y in a..b, deriv f y` equals `f b - f a`. -/ -theorem integral_deriv_eq_sub (hderiv : ∀ x ∈ interval a b, differentiable_at ℝ f x) +theorem integral_deriv_eq_sub (hderiv : ∀ x ∈ uIcc a b, differentiable_at ℝ f x) (hint : interval_integrable (deriv f) volume a b) : ∫ y in a..b, deriv f y = f b - f a := integral_eq_sub_of_has_deriv_at (λ x hx, (hderiv x hx).has_deriv_at) hint theorem integral_deriv_eq_sub' (f) (hderiv : deriv f = f') - (hdiff : ∀ x ∈ interval a b, differentiable_at ℝ f x) - (hcont : continuous_on f' (interval a b)) : + (hdiff : ∀ x ∈ uIcc a b, differentiable_at ℝ f x) + (hcont : continuous_on f' (uIcc a b)) : ∫ y in a..b, f' y = f b - f a := begin rw [← hderiv, integral_deriv_eq_sub hdiff], @@ -2470,16 +2470,16 @@ integrable_on_deriv_right_of_nonneg hab hcont (λ x hx, (hderiv x hx).has_deriv_ /-- When the derivative of a function is nonnegative, then it is automatically integrable, interval version. -/ -theorem interval_integrable_deriv_of_nonneg (hcont : continuous_on g (interval a b)) +theorem interval_integrable_deriv_of_nonneg (hcont : continuous_on g (uIcc a b)) (hderiv : ∀ x ∈ Ioo (min a b) (max a b), has_deriv_at g (g' x) x) (hpos : ∀ x ∈ Ioo (min a b) (max a b), 0 ≤ g' x) : interval_integrable g' volume a b := begin cases le_total a b with hab hab, - { simp only [interval_of_le, min_eq_left, max_eq_right, hab, interval_integrable, + { simp only [uIcc_of_le, min_eq_left, max_eq_right, hab, interval_integrable, hab, Ioc_eq_empty_of_le, integrable_on_empty, and_true] at hcont hderiv hpos ⊢, exact integrable_on_deriv_of_nonneg hab hcont hderiv hpos, }, - { simp only [interval_of_ge, min_eq_right, max_eq_left, hab, interval_integrable, + { simp only [uIcc_of_ge, min_eq_right, max_eq_left, hab, interval_integrable, Ioc_eq_empty_of_le, integrable_on_empty, true_and] at hcont hderiv hpos ⊢, exact integrable_on_deriv_of_nonneg hab hcont hderiv hpos } end @@ -2493,8 +2493,8 @@ section parts variables [normed_ring A] [normed_algebra ℝ A] [complete_space A] theorem integral_deriv_mul_eq_sub {u v u' v' : ℝ → A} - (hu : ∀ x ∈ interval a b, has_deriv_at u (u' x) x) - (hv : ∀ x ∈ interval a b, has_deriv_at v (v' x) x) + (hu : ∀ x ∈ uIcc a b, has_deriv_at u (u' x) x) + (hv : ∀ x ∈ uIcc a b, has_deriv_at v (v' x) x) (hu' : interval_integrable u' volume a b) (hv' : interval_integrable v' volume a b) : ∫ x in a..b, u' x * v x + u x * v' x = u b * v b - u a * v a := integral_eq_sub_of_has_deriv_at (λ x hx, (hu x hx).mul (hv x hx)) $ @@ -2502,8 +2502,8 @@ integral_eq_sub_of_has_deriv_at (λ x hx, (hu x hx).mul (hv x hx)) $ (hv'.continuous_on_mul ((has_deriv_at.continuous_on hu))) theorem integral_mul_deriv_eq_deriv_mul {u v u' v' : ℝ → A} - (hu : ∀ x ∈ interval a b, has_deriv_at u (u' x) x) - (hv : ∀ x ∈ interval a b, has_deriv_at v (v' x) x) + (hu : ∀ x ∈ uIcc a b, has_deriv_at u (u' x) x) + (hv : ∀ x ∈ uIcc a b, has_deriv_at v (v' x) x) (hu' : interval_integrable u' volume a b) (hv' : interval_integrable v' volume a b) : ∫ x in a..b, u x * v' x = u b * v b - u a * v a - ∫ x in a..b, u' x * v x := begin @@ -2536,27 +2536,27 @@ theorem integral_comp_smul_deriv''' {f f' : ℝ → ℝ} {g : ℝ → E} (hg2 : integrable_on (λ x, f'(x) • (g ∘ f) x) [a, b]) : ∫ x in a..b, f' x • (g ∘ f) x = ∫ u in f a..f b, g u := begin - rw [hf.image_interval, ←interval_integrable_iff'] at hg1, + rw [hf.image_uIcc, ←interval_integrable_iff'] at hg1, have h_cont : continuous_on (λ u, ∫ t in f a..f u, g t) [a, b], { refine (continuous_on_primitive_interval' hg1 _).comp hf _, - { rw ← hf.image_interval, exact mem_image_of_mem f left_mem_interval }, - { rw ← hf.image_interval, exact maps_to_image _ _ } }, + { rw ← hf.image_uIcc, exact mem_image_of_mem f left_mem_uIcc }, + { rw ← hf.image_uIcc, exact maps_to_image _ _ } }, have h_der : ∀ x ∈ Ioo (min a b) (max a b), has_deriv_within_at (λ u, ∫ t in f a..f u, g t) (f' x • ((g ∘ f) x)) (Ioi x) x, { intros x hx, obtain ⟨c, hc⟩ := nonempty_Ioo.mpr hx.1, obtain ⟨d, hd⟩ := nonempty_Ioo.mpr hx.2, have cdsub : [c, d] ⊆ Ioo (min a b) (max a b), - { rw interval_of_le (hc.2.trans hd.1).le, exact Icc_subset_Ioo hc.1 hd.2 }, + { rw uIcc_of_le (hc.2.trans hd.1).le, exact Icc_subset_Ioo hc.1 hd.2 }, replace hg_cont := hg_cont.mono (image_subset f cdsub), let J := [Inf (f '' [c, d]), Sup (f '' [c, d])], - have hJ : f '' [c, d] = J := (hf.mono (cdsub.trans Ioo_subset_Icc_self)).image_interval, + have hJ : f '' [c, d] = J := (hf.mono (cdsub.trans Ioo_subset_Icc_self)).image_uIcc, rw hJ at hg_cont, - have h2x : f x ∈ J, { rw ←hJ, exact mem_image_of_mem _ (mem_interval_of_le hc.2.le hd.1.le), }, + have h2x : f x ∈ J, { rw ←hJ, exact mem_image_of_mem _ (mem_uIcc_of_le hc.2.le hd.1.le), }, have h2g : interval_integrable g volume (f a) (f x), { refine hg1.mono_set _, - rw ←hf.image_interval, - exact hf.surj_on_interval left_mem_interval (Ioo_subset_Icc_self hx) }, + rw ←hf.image_uIcc, + exact hf.surj_on_uIcc left_mem_uIcc (Ioo_subset_Icc_self hx) }, have h3g := hg_cont.strongly_measurable_at_filter_nhds_within measurable_set_Icc (f x), haveI : fact (f x ∈ J) := ⟨h2x⟩, have : has_deriv_within_at (λ u, ∫ x in f a..u, g x) (g (f x)) J (f x) := @@ -2564,7 +2564,7 @@ begin refine (this.scomp x ((hff' x hx).Ioo_of_Ioi hd.1) _).Ioi_of_Ioo hd.1, rw ←hJ, refine (maps_to_image _ _).mono _ subset.rfl, - exact Ioo_subset_Icc_self.trans ((Icc_subset_Icc_left hc.2.le).trans Icc_subset_interval) }, + exact Ioo_subset_Icc_self.trans ((Icc_subset_Icc_left hc.2.le).trans Icc_subset_uIcc) }, rw ←interval_integrable_iff' at hg2, simp_rw [integral_eq_sub_of_has_deriv_right h_cont h_der hg2, integral_same, sub_zero], end @@ -2584,7 +2584,7 @@ begin refine integral_comp_smul_deriv''' hf hff' (hg.mono $ image_subset _ Ioo_subset_Icc_self) _ (hf'.smul (hg.comp hf $ subset_preimage_image f _)).integrable_on_Icc, - rw hf.image_interval at hg ⊢, + rw hf.image_uIcc at hg ⊢, exact hg.integrable_on_Icc, end @@ -2596,8 +2596,8 @@ Compared to `interval_integral.integral_comp_smul_deriv` we only require that `g `f '' [a, b]`. -/ theorem integral_comp_smul_deriv' {f f' : ℝ → ℝ} {g : ℝ → E} - (h : ∀ x ∈ interval a b, has_deriv_at f (f' x) x) - (h' : continuous_on f' (interval a b)) (hg : continuous_on g (f '' [a, b])) : + (h : ∀ x ∈ uIcc a b, has_deriv_at f (f' x) x) + (h' : continuous_on f' (uIcc a b)) (hg : continuous_on g (f '' [a, b])) : ∫ x in a..b, f' x • (g ∘ f) x = ∫ x in f a..f b, g x := integral_comp_smul_deriv'' (λ x hx, (h x hx).continuous_at.continuous_within_at) (λ x hx, (h x $ Ioo_subset_Icc_self hx).has_deriv_within_at) h' hg @@ -2608,8 +2608,8 @@ and `g` is continuous, then we can substitute `u = f x` to get `∫ x in a..b, f' x • (g ∘ f) x = ∫ u in f a..f b, g u`. -/ theorem integral_comp_smul_deriv {f f' : ℝ → ℝ} {g : ℝ → E} - (h : ∀ x ∈ interval a b, has_deriv_at f (f' x) x) - (h' : continuous_on f' (interval a b)) (hg : continuous g) : + (h : ∀ x ∈ uIcc a b, has_deriv_at f (f' x) x) + (h' : continuous_on f' (uIcc a b)) (hg : continuous g) : ∫ x in a..b, f' x • (g ∘ f) x = ∫ x in f a..f b, g x := integral_comp_smul_deriv' h h' hg.continuous_on @@ -2624,13 +2624,13 @@ theorem integral_deriv_comp_smul_deriv' {f f' : ℝ → ℝ} {g g' : ℝ → E} begin rw [integral_comp_smul_deriv'' hf hff' hf' hg', integral_eq_sub_of_has_deriv_right hg hgg' (hg'.mono _).interval_integrable], - exact intermediate_value_interval hf + exact intermediate_value_uIcc hf end theorem integral_deriv_comp_smul_deriv {f f' : ℝ → ℝ} {g g' : ℝ → E} - (hf : ∀ x ∈ interval a b, has_deriv_at f (f' x) x) - (hg : ∀ x ∈ interval a b, has_deriv_at g (g' (f x)) (f x)) - (hf' : continuous_on f' (interval a b)) (hg' : continuous g') : + (hf : ∀ x ∈ uIcc a b, has_deriv_at f (f' x) x) + (hg : ∀ x ∈ uIcc a b, has_deriv_at g (g' (f x)) (f x)) + (hf' : continuous_on f' (uIcc a b)) (hg' : continuous g') : ∫ x in a..b, f' x • (g' ∘ f) x = (g ∘ f) b - (g ∘ f) a := integral_eq_sub_of_has_deriv_at (λ x hx, (hg x hx).scomp x $ hf x hx) (hf'.smul (hg'.comp_continuous_on $ has_deriv_at.continuous_on hf)).interval_integrable @@ -2677,8 +2677,8 @@ Compared to `interval_integral.integral_comp_mul_deriv` we only require that `g` `f '' [a, b]`. -/ theorem integral_comp_mul_deriv' {f f' g : ℝ → ℝ} - (h : ∀ x ∈ interval a b, has_deriv_at f (f' x) x) - (h' : continuous_on f' (interval a b)) (hg : continuous_on g (f '' [a, b])) : + (h : ∀ x ∈ uIcc a b, has_deriv_at f (f' x) x) + (h' : continuous_on f' (uIcc a b)) (hg : continuous_on g (f '' [a, b])) : ∫ x in a..b, (g ∘ f) x * f' x = ∫ x in f a..f b, g x := by simpa [mul_comm] using integral_comp_smul_deriv' h h' hg @@ -2688,8 +2688,8 @@ and `g` is continuous, then we can substitute `u = f x` to get `∫ x in a..b, (g ∘ f) x * f' x = ∫ u in f a..f b, g u`. -/ theorem integral_comp_mul_deriv {f f' g : ℝ → ℝ} - (h : ∀ x ∈ interval a b, has_deriv_at f (f' x) x) - (h' : continuous_on f' (interval a b)) (hg : continuous g) : + (h : ∀ x ∈ uIcc a b, has_deriv_at f (f' x) x) + (h' : continuous_on f' (uIcc a b)) (hg : continuous g) : ∫ x in a..b, (g ∘ f) x * f' x = ∫ x in f a..f b, g x := integral_comp_mul_deriv' h h' hg.continuous_on @@ -2704,9 +2704,9 @@ theorem integral_deriv_comp_mul_deriv' {f f' g g' : ℝ → ℝ} by simpa [mul_comm] using integral_deriv_comp_smul_deriv' hf hff' hf' hg hgg' hg' theorem integral_deriv_comp_mul_deriv {f f' g g' : ℝ → ℝ} - (hf : ∀ x ∈ interval a b, has_deriv_at f (f' x) x) - (hg : ∀ x ∈ interval a b, has_deriv_at g (g' (f x)) (f x)) - (hf' : continuous_on f' (interval a b)) (hg' : continuous g') : + (hf : ∀ x ∈ uIcc a b, has_deriv_at f (f' x) x) + (hg : ∀ x ∈ uIcc a b, has_deriv_at g (g' (f x)) (f x)) + (hf' : continuous_on f' (uIcc a b)) (hg' : continuous g') : ∫ x in a..b, (g' ∘ f) x * f' x = (g ∘ f) b - (g ∘ f) a := by simpa [mul_comm] using integral_deriv_comp_smul_deriv hf hg hf' hg' diff --git a/src/measure_theory/integral/layercake.lean b/src/measure_theory/integral/layercake.lean index d4975c79d6bd8..f95873332e603 100644 --- a/src/measure_theory/integral/layercake.lean +++ b/src/measure_theory/integral/layercake.lean @@ -181,7 +181,7 @@ begin have eq₂ : ∀ ω, ∫ t in 0..f ω, g t = ∫ t in 0..f ω, G t, { refine λ ω, interval_integral.integral_congr_ae _, have fω_nn : 0 ≤ f ω := f_nn ω, - rw [interval_oc_of_le fω_nn, + rw [uIoc_of_le fω_nn, ← ae_restrict_iff' (measurable_set_Ioc : measurable_set (Ioc (0 : ℝ) (f ω)))], exact g_eq_G_on (f ω), }, simp_rw [eq₁, eq₂], diff --git a/src/measure_theory/measure/ae_measurable.lean b/src/measure_theory/measure/ae_measurable.lean index 5a97a8502d46f..c947c59a067b5 100644 --- a/src/measure_theory/measure/ae_measurable.lean +++ b/src/measure_theory/measure/ae_measurable.lean @@ -208,10 +208,10 @@ let ⟨g, hgm, hg⟩ := h in hgm.null_measurable.congr hg.symm end ae_measurable -lemma ae_measurable_interval_oc_iff [linear_order α] {f : α → β} {a b : α} : +lemma ae_measurable_uIoc_iff [linear_order α] {f : α → β} {a b : α} : (ae_measurable f $ μ.restrict $ Ι a b) ↔ (ae_measurable f $ μ.restrict $ Ioc a b) ∧ (ae_measurable f $ μ.restrict $ Ioc b a) := -by rw [interval_oc_eq_union, ae_measurable_union_iff] +by rw [uIoc_eq_union, ae_measurable_union_iff] lemma ae_measurable_iff_measurable [μ.is_complete] : ae_measurable f μ ↔ measurable f := diff --git a/src/measure_theory/measure/lebesgue.lean b/src/measure_theory/measure/lebesgue.lean index 2d20d0e8656bf..b69c07450b49a 100644 --- a/src/measure_theory/measure/lebesgue.lean +++ b/src/measure_theory/measure/lebesgue.lean @@ -112,7 +112,7 @@ end instance has_no_atoms_volume : has_no_atoms (volume : measure ℝ) := ⟨λ x, volume_singleton⟩ -@[simp] lemma volume_interval {a b : ℝ} : volume (interval a b) = of_real (|b - a|) := +@[simp] lemma volume_interval {a b : ℝ} : volume (uIcc a b) = of_real (|b - a|) := by rw [←Icc_min_max, volume_Icc, max_sub_min_eq_abs] @[simp] lemma volume_Ioi {a : ℝ} : volume (Ioi a) = ∞ := diff --git a/src/measure_theory/measure/measure_space.lean b/src/measure_theory/measure/measure_space.lean index 75c82f98d98a1..dc31c755e6e02 100644 --- a/src/measure_theory/measure/measure_space.lean +++ b/src/measure_theory/measure/measure_space.lean @@ -104,10 +104,10 @@ instance ae_is_measurably_generated : is_measurably_generated μ.ae := ⟨λ s hs, let ⟨t, hst, htm, htμ⟩ := exists_measurable_superset_of_null hs in ⟨tᶜ, compl_mem_ae_iff.2 htμ, htm.compl, compl_subset_comm.1 hst⟩⟩ -/-- See also `measure_theory.ae_restrict_interval_oc_iff`. -/ -lemma ae_interval_oc_iff [linear_order α] {a b : α} {P : α → Prop} : +/-- See also `measure_theory.ae_restrict_uIoc_iff`. -/ +lemma ae_uIoc_iff [linear_order α] {a b : α} {P : α → Prop} : (∀ᵐ x ∂μ, x ∈ Ι a b → P x) ↔ (∀ᵐ x ∂μ, x ∈ Ioc a b → P x) ∧ (∀ᵐ x ∂μ, x ∈ Ioc b a → P x) := -by simp only [interval_oc_eq_union, mem_union, or_imp_distrib, eventually_and] +by simp only [uIoc_eq_union, mem_union, or_imp_distrib, eventually_and] lemma measure_union (hd : disjoint s₁ s₂) (h : measurable_set s₂) : μ (s₁ ∪ s₂) = μ s₁ + μ s₂ := @@ -2381,15 +2381,15 @@ lemma ae_eq_restrict_bUnion_finset_iff (s : ι → set α) (t : finset ι) (f g f =ᵐ[μ.restrict (⋃ i ∈ t, s i)] g ↔ ∀ i ∈ t, f =ᵐ[μ.restrict (s i)] g := ae_eq_restrict_bUnion_iff s t.countable_to_set f g -lemma ae_restrict_interval_oc_eq [linear_order α] (a b : α) : +lemma ae_restrict_uIoc_eq [linear_order α] (a b : α) : (μ.restrict (Ι a b)).ae = (μ.restrict (Ioc a b)).ae ⊔ (μ.restrict (Ioc b a)).ae := -by simp only [interval_oc_eq_union, ae_restrict_union_eq] +by simp only [uIoc_eq_union, ae_restrict_union_eq] -/-- See also `measure_theory.ae_interval_oc_iff`. -/ -lemma ae_restrict_interval_oc_iff [linear_order α] {a b : α} {P : α → Prop} : +/-- See also `measure_theory.ae_uIoc_iff`. -/ +lemma ae_restrict_uIoc_iff [linear_order α] {a b : α} {P : α → Prop} : (∀ᵐ x ∂μ.restrict (Ι a b), P x) ↔ (∀ᵐ x ∂μ.restrict (Ioc a b), P x) ∧ (∀ᵐ x ∂μ.restrict (Ioc b a), P x) := -by rw [ae_restrict_interval_oc_eq, eventually_sup] +by rw [ae_restrict_uIoc_eq, eventually_sup] lemma ae_restrict_iff {p : α → Prop} (hp : measurable_set {x | p x}) : (∀ᵐ x ∂(μ.restrict s), p x) ↔ ∀ᵐ x ∂μ, x ∈ s → p x := @@ -2900,7 +2900,7 @@ end open_locale interval -lemma interval_oc_ae_eq_interval [linear_order α] {a b : α} : Ι a b =ᵐ[μ] [a, b] := Ioc_ae_eq_Icc +lemma uIoc_ae_eq_interval [linear_order α] {a b : α} : Ι a b =ᵐ[μ] [a, b] := Ioc_ae_eq_Icc end no_atoms diff --git a/src/order/filter/interval.lean b/src/order/filter/interval.lean index 605bfcd00266e..ce637e7d62fdb 100644 --- a/src/order/filter/interval.lean +++ b/src/order/filter/interval.lean @@ -198,25 +198,25 @@ section linear_order variables [linear_order α] -instance tendsto_Icc_interval_interval {a b : α} : tendsto_Ixx_class Icc (𝓟 [a, b]) (𝓟 [a, b]) := +instance tendsto_Icc_uIcc_uIcc {a b : α} : tendsto_Ixx_class Icc (𝓟 [a, b]) (𝓟 [a, b]) := filter.tendsto_Icc_Icc_Icc -instance tendsto_Ioc_interval_interval {a b : α} : tendsto_Ixx_class Ioc (𝓟 [a, b]) (𝓟 [a, b]) := +instance tendsto_Ioc_uIcc_uIcc {a b : α} : tendsto_Ixx_class Ioc (𝓟 [a, b]) (𝓟 [a, b]) := filter.tendsto_Ioc_Icc_Icc -instance tendsto_interval_of_Icc {l : filter α} [tendsto_Ixx_class Icc l l] : - tendsto_Ixx_class interval l l := +instance tendsto_uIcc_of_Icc {l : filter α} [tendsto_Ixx_class Icc l l] : + tendsto_Ixx_class uIcc l l := begin refine ⟨λ s hs, mem_map.2 $ mem_prod_self_iff.2 _⟩, obtain ⟨t, htl, hts⟩ : ∃ t ∈ l, ∀ p ∈ (t : set α) ×ˢ t, Icc (p : α × α).1 p.2 ∈ s, from mem_prod_self_iff.1 (mem_map.1 (tendsto_fst.Icc tendsto_snd hs)), refine ⟨t, htl, λ p hp, _⟩, cases le_total p.1 p.2, - { rw [mem_preimage, interval_of_le h], exact hts p hp }, - { rw [mem_preimage, interval_of_ge h], exact hts ⟨p.2, p.1⟩ ⟨hp.2, hp.1⟩ } + { rw [mem_preimage, uIcc_of_le h], exact hts p hp }, + { rw [mem_preimage, uIcc_of_ge h], exact hts ⟨p.2, p.1⟩ ⟨hp.2, hp.1⟩ } end -lemma tendsto.interval {l : filter α} [tendsto_Ixx_class Icc l l] {f g : β → α} {lb : filter β} +lemma tendsto.uIcc {l : filter α} [tendsto_Ixx_class Icc l l] {f g : β → α} {lb : filter β} (hf : tendsto f lb l) (hg : tendsto g lb l) : tendsto (λ x, [f x, g x]) lb l.small_sets := tendsto_Ixx_class.tendsto_Ixx.comp $ hf.prod_mk hg diff --git a/src/order/locally_finite.lean b/src/order/locally_finite.lean index 037ce5d44d580..e1e3238b53a60 100644 --- a/src/order/locally_finite.lean +++ b/src/order/locally_finite.lean @@ -30,7 +30,7 @@ In a `locally_finite_order`, * `finset.Ico`: Closed-open interval as a finset. * `finset.Ioc`: Open-closed interval as a finset. * `finset.Ioo`: Open-open interval as a finset. -* `finset.interval`: Unordered closed interval as a finset. +* `finset.uIcc`: Unordered closed interval as a finset. * `multiset.Icc`: Closed-closed interval as a multiset. * `multiset.Ico`: Closed-open interval as a multiset. * `multiset.Ioc`: Open-closed interval as a multiset. @@ -367,17 +367,17 @@ end preorder section lattice variables [lattice α] [locally_finite_order α] {a b x : α} -/-- `finset.interval a b` is the set of elements lying between `a` and `b`, with `a` and `b` -included. Note that we define it more generally in a lattice as `finset.Icc (a ⊓ b) (a ⊔ b)`. In a -product type, `finset.interval` corresponds to the bounding box of the two elements. -/ -def interval (a b : α) : finset α := Icc (a ⊓ b) (a ⊔ b) +/-- `finset.uIcc a b` is the set of elements lying between `a` and `b`, with `a` and `b` included. +Note that we define it more generally in a lattice as `finset.Icc (a ⊓ b) (a ⊔ b)`. In a +product type, `finset.uIcc` corresponds to the bounding box of the two elements. -/ +def uIcc (a b : α) : finset α := Icc (a ⊓ b) (a ⊔ b) -localized "notation (name := finset.interval) `[`a `, ` b `]` := finset.interval a b" +localized "notation (name := finset.uIcc) `[`a `, ` b `]` := finset.uIcc a b" in finset_interval -@[simp] lemma mem_interval : x ∈ interval a b ↔ a ⊓ b ≤ x ∧ x ≤ a ⊔ b := mem_Icc +@[simp] lemma mem_uIcc : x ∈ uIcc a b ↔ a ⊓ b ≤ x ∧ x ≤ a ⊔ b := mem_Icc -@[simp, norm_cast] lemma coe_interval (a b : α) : ([a, b] : set α) = set.interval a b := coe_Icc _ _ +@[simp, norm_cast] lemma coe_uIcc (a b : α) : ([a, b] : set α) = set.uIcc a b := coe_Icc _ _ end lattice end finset @@ -718,17 +718,17 @@ end preorder namespace prod variables [lattice α] [lattice β] -lemma interval_eq [locally_finite_order α] [locally_finite_order β] +lemma uIcc_eq [locally_finite_order α] [locally_finite_order β] [decidable_rel ((≤) : α × β → α × β → Prop)] (p q : α × β) : - finset.interval p q = finset.interval p.1 q.1 ×ˢ finset.interval p.2 q.2 := rfl + finset.uIcc p q = finset.uIcc p.1 q.1 ×ˢ finset.uIcc p.2 q.2 := rfl -@[simp] lemma interval_mk_mk [locally_finite_order α] [locally_finite_order β] +@[simp] lemma uIcc_mk_mk [locally_finite_order α] [locally_finite_order β] [decidable_rel ((≤) : α × β → α × β → Prop)] (a₁ a₂ : α) (b₁ b₂ : β) : - finset.interval (a₁, b₁) (a₂, b₂) = finset.interval a₁ a₂ ×ˢ finset.interval b₁ b₂ := rfl + finset.uIcc (a₁, b₁) (a₂, b₂) = finset.uIcc a₁ a₂ ×ˢ finset.uIcc b₁ b₂ := rfl -lemma card_interval [locally_finite_order α] [locally_finite_order β] +lemma card_uIcc [locally_finite_order α] [locally_finite_order β] [decidable_rel ((≤) : α × β → α × β → Prop)] (p q : α × β) : - (finset.interval p q).card = (finset.interval p.1 q.1).card * (finset.interval p.2 q.2).card := + (finset.uIcc p q).card = (finset.uIcc p.1 q.1).card * (finset.uIcc p.2 q.2).card := prod.card_Icc _ _ end prod diff --git a/src/topology/algebra/order/compact.lean b/src/topology/algebra/order/compact.lean index 2286eeec194fd..c2aa39290434e 100644 --- a/src/topology/algebra/order/compact.lean +++ b/src/topology/algebra/order/compact.lean @@ -108,8 +108,8 @@ instance {α β : Type*} [preorder α] [topological_space α] [compact_Icc_space ⟨λ a b, (Icc_prod_eq a b).symm ▸ is_compact_Icc.prod is_compact_Icc⟩ /-- An unordered closed interval is compact. -/ -lemma is_compact_interval {α : Type*} [linear_order α] [topological_space α] [compact_Icc_space α] - {a b : α} : is_compact (interval a b) := +lemma is_compact_uIcc {α : Type*} [linear_order α] [topological_space α] [compact_Icc_space α] + {a b : α} : is_compact (uIcc a b) := is_compact_Icc /-- A complete linear order is a compact space. @@ -356,19 +356,19 @@ lemma image_Icc (hab : a ≤ b) (h : continuous_on f $ Icc a b) : eq_Icc_of_connected_compact ⟨(nonempty_Icc.2 hab).image f, is_preconnected_Icc.image f h⟩ (is_compact_Icc.image_of_continuous_on h) -lemma image_interval_eq_Icc (h : continuous_on f $ [a, b]) : +lemma image_uIcc_eq_Icc (h : continuous_on f $ [a, b]) : f '' [a, b] = Icc (Inf (f '' [a, b])) (Sup (f '' [a, b])) := begin cases le_total a b with h2 h2, - { simp_rw [interval_of_le h2] at h ⊢, exact h.image_Icc h2 }, - { simp_rw [interval_of_ge h2] at h ⊢, exact h.image_Icc h2 }, + { simp_rw [uIcc_of_le h2] at h ⊢, exact h.image_Icc h2 }, + { simp_rw [uIcc_of_ge h2] at h ⊢, exact h.image_Icc h2 }, end -lemma image_interval (h : continuous_on f $ [a, b]) : +lemma image_uIcc (h : continuous_on f $ [a, b]) : f '' [a, b] = [Inf (f '' [a, b]), Sup (f '' [a, b])] := begin - refine h.image_interval_eq_Icc.trans (interval_of_le _).symm, - refine cInf_le_cSup _ _ (nonempty_interval.image _); rw h.image_interval_eq_Icc, + refine h.image_uIcc_eq_Icc.trans (uIcc_of_le _).symm, + refine cInf_le_cSup _ _ (nonempty_uIcc.image _); rw h.image_uIcc_eq_Icc, exacts [bdd_below_Icc, bdd_above_Icc] end diff --git a/src/topology/algebra/order/intermediate_value.lean b/src/topology/algebra/order/intermediate_value.lean index 7a07491077eb7..77dad319a8919 100644 --- a/src/topology/algebra/order/intermediate_value.lean +++ b/src/topology/algebra/order/intermediate_value.lean @@ -418,12 +418,12 @@ begin exact is_preconnected_Icc_aux y x t s h ht hs hab hy hx, }, end -lemma is_preconnected_interval : is_preconnected (interval a b) := is_preconnected_Icc +lemma is_preconnected_uIcc : is_preconnected (uIcc a b) := is_preconnected_Icc lemma set.ord_connected.is_preconnected {s : set α} (h : s.ord_connected) : is_preconnected s := -is_preconnected_of_forall_pair $ λ x hx y hy, ⟨interval x y, h.interval_subset hx hy, - left_mem_interval, right_mem_interval, is_preconnected_interval⟩ +is_preconnected_of_forall_pair $ λ x hx y hy, ⟨uIcc x y, h.uIcc_subset hx hy, + left_mem_uIcc, right_mem_uIcc, is_preconnected_uIcc⟩ lemma is_preconnected_iff_ord_connected {s : set α} : is_preconnected s ↔ ord_connected s := @@ -504,9 +504,9 @@ lemma intermediate_value_Icc' {a b : α} (hab : a ≤ b) {f : α → δ} (hf : c is_preconnected_Icc.intermediate_value (right_mem_Icc.2 hab) (left_mem_Icc.2 hab) hf /-- **Intermediate Value Theorem** for continuous functions on closed intervals, unordered case. -/ -lemma intermediate_value_interval {a b : α} {f : α → δ} (hf : continuous_on f (interval a b)) : - interval (f a) (f b) ⊆ f '' interval a b := -by cases le_total (f a) (f b); simp [*, is_preconnected_interval.intermediate_value] +lemma intermediate_value_uIcc {a b : α} {f : α → δ} (hf : continuous_on f (uIcc a b)) : + uIcc (f a) (f b) ⊆ f '' uIcc a b := +by cases le_total (f a) (f b); simp [*, is_preconnected_uIcc.intermediate_value] lemma intermediate_value_Ico {a b : α} (hab : a ≤ b) {f : α → δ} (hf : continuous_on f (Icc a b)) : Ico (f a) (f b) ⊆ f '' (Ico a b) := @@ -563,9 +563,9 @@ hs.is_preconnected.intermediate_value ha hb hf /-- **Intermediate value theorem**: if `f` is continuous on an order-connected set `s` and `a`, `b` are two points of this set, then `f` sends `s` to a superset of `[f x, f y]`. -/ -lemma continuous_on.surj_on_interval {s : set α} [hs : ord_connected s] {f : α → δ} +lemma continuous_on.surj_on_uIcc {s : set α} [hs : ord_connected s] {f : α → δ} (hf : continuous_on f s) {a b : α} (ha : a ∈ s) (hb : b ∈ s) : - surj_on f s (interval (f a) (f b)) := + surj_on f s (uIcc (f a) (f b)) := by cases le_total (f a) (f b) with hab hab; simp [hf.surj_on_Icc, *] /-- A continuous function which tendsto `at_top` `at_top` and to `at_bot` `at_bot` is surjective. -/ diff --git a/src/topology/algebra/order/t5.lean b/src/topology/algebra/order/t5.lean index 504e4106a808b..f2a48d583681c 100644 --- a/src/topology/algebra/order/t5.lean +++ b/src/topology/algebra/order/t5.lean @@ -47,19 +47,19 @@ begin replace hac : a < c := hac.lt_of_ne (ne.symm $ ne_of_mem_of_not_mem hc $ disjoint_left.1 (disjoint_left_ord_separating_set.mono_right ord_connected_section_subset) ha), refine mem_of_superset (Ico_mem_nhds_within_Ici (left_mem_Ico.2 hac)) (λ x hx hx', _), - refine hx.2.ne (eq_of_mem_ord_connected_section_of_interval_subset hx' hc _), + refine hx.2.ne (eq_of_mem_ord_connected_section_of_uIcc_subset hx' hc _), refine subset_inter (subset_Union₂_of_subset a ha _) _, - { exact ord_connected.interval_subset infer_instance (hsub' ⟨hx.1, hx.2.le.trans hcb⟩) + { exact ord_connected.uIcc_subset infer_instance (hsub' ⟨hx.1, hx.2.le.trans hcb⟩) (hsub' ⟨hac.le, hcb⟩) }, { rcases mem_Union₂.1 (ord_connected_section_subset hx').2 with ⟨y, hyt, hxy⟩, - refine subset_Union₂_of_subset y hyt (ord_connected.interval_subset infer_instance hxy _), - refine subset_ord_connected_component left_mem_interval hxy _, + refine subset_Union₂_of_subset y hyt (ord_connected.uIcc_subset infer_instance hxy _), + refine subset_ord_connected_component left_mem_uIcc hxy _, suffices : c < y, - { rw [interval_of_ge (hx.2.trans this).le], + { rw [uIcc_of_ge (hx.2.trans this).le], exact ⟨hx.2.le, this.le⟩ }, refine lt_of_not_le (λ hyc, _), have hya : y < a, from not_le.1 (λ hay, hsub ⟨hay, hyc.trans hcb⟩ hyt), - exact hxy (Icc_subset_interval ⟨hya.le, hx.1⟩) ha } } + exact hxy (Icc_subset_uIcc ⟨hya.le, hx.1⟩) ha } } end lemma compl_section_ord_separating_set_mem_nhds_within_Iic (hd : disjoint s (closure t)) diff --git a/src/topology/metric_space/basic.lean b/src/topology/metric_space/basic.lean index c219fc00273a7..c00763637c991 100644 --- a/src/topology/metric_space/basic.lean +++ b/src/topology/metric_space/basic.lean @@ -1242,22 +1242,22 @@ theorem real.nndist_eq' (x y : ℝ) : nndist x y = real.nnabs (y - x) := nndist_ theorem real.dist_0_eq_abs (x : ℝ) : dist x 0 = |x| := by simp [real.dist_eq] -theorem real.dist_left_le_of_mem_interval {x y z : ℝ} (h : y ∈ interval x z) : +theorem real.dist_left_le_of_mem_uIcc {x y z : ℝ} (h : y ∈ uIcc x z) : dist x y ≤ dist x z := -by simpa only [dist_comm x] using abs_sub_left_of_mem_interval h +by simpa only [dist_comm x] using abs_sub_left_of_mem_uIcc h -theorem real.dist_right_le_of_mem_interval {x y z : ℝ} (h : y ∈ interval x z) : +theorem real.dist_right_le_of_mem_uIcc {x y z : ℝ} (h : y ∈ uIcc x z) : dist y z ≤ dist x z := -by simpa only [dist_comm _ z] using abs_sub_right_of_mem_interval h +by simpa only [dist_comm _ z] using abs_sub_right_of_mem_uIcc h -theorem real.dist_le_of_mem_interval {x y x' y' : ℝ} (hx : x ∈ interval x' y') - (hy : y ∈ interval x' y') : dist x y ≤ dist x' y' := -abs_sub_le_of_subinterval $ interval_subset_interval (by rwa interval_swap) (by rwa interval_swap) +theorem real.dist_le_of_mem_uIcc {x y x' y' : ℝ} (hx : x ∈ uIcc x' y') + (hy : y ∈ uIcc x' y') : dist x y ≤ dist x' y' := +abs_sub_le_of_uIcc_subset_uIcc $ uIcc_subset_uIcc (by rwa uIcc_comm) (by rwa uIcc_comm) theorem real.dist_le_of_mem_Icc {x y x' y' : ℝ} (hx : x ∈ Icc x' y') (hy : y ∈ Icc x' y') : dist x y ≤ y' - x' := by simpa only [real.dist_eq, abs_of_nonpos (sub_nonpos.2 $ hx.1.trans hx.2), neg_sub] - using real.dist_le_of_mem_interval (Icc_subset_interval hx) (Icc_subset_interval hy) + using real.dist_le_of_mem_uIcc (Icc_subset_uIcc hx) (Icc_subset_uIcc hy) theorem real.dist_le_of_mem_Icc_01 {x y : ℝ} (hx : x ∈ Icc (0:ℝ) 1) (hy : y ∈ Icc (0:ℝ) 1) : dist x y ≤ 1 := @@ -1913,8 +1913,8 @@ by simp only [dist_nndist, fin.nndist_insert_nth_insert_nth, nnreal.coe_max] lemma real.dist_le_of_mem_pi_Icc {x y x' y' : β → ℝ} (hx : x ∈ Icc x' y') (hy : y ∈ Icc x' y') : dist x y ≤ dist x' y' := begin - refine (dist_pi_le_iff dist_nonneg).2 (λ b, (real.dist_le_of_mem_interval _ _).trans - (dist_le_pi_dist _ _ b)); refine Icc_subset_interval _, + refine (dist_pi_le_iff dist_nonneg).2 (λ b, (real.dist_le_of_mem_uIcc _ _).trans + (dist_le_pi_dist _ _ b)); refine Icc_subset_uIcc _, exacts [⟨hx.1 _, hx.2 _⟩, ⟨hy.1 _, hy.2 _⟩] end From 7e7aaccf9b0182576cabdde36cf1b5ad3585b70d Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 12 Jan 2023 20:46:50 +0000 Subject: [PATCH 0224/1291] chore(*): remove some defeq abuse of `^` (#18136) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit There are various places where the proofs relied on the defeq of `^`, namely that * `rfl : x ^ 0 = 1` * `rfl : x ^ (n : ℕ) = x ^ (↑n : ℤ)` * `rfl : x ^ (n.succ : ℕ) = x * x ^ n` While exploiting defeq is not always a bad thing, in some of these cases it was clear that this wasn't intentional, and the author had not noticed that they were working with integer instead of natural powers. This change is motivated by two possible refactors: * Changing the definition of `monoid.npow_rec` to match `function.iterate` * Changing the definition of `real.has_pow` to be the elementwise power on the cauchy sequence (#8146) --- src/analysis/special_functions/bernstein.lean | 3 ++- src/analysis/special_functions/integrals.lean | 2 +- src/analysis/special_functions/pow.lean | 8 ++++---- src/analysis/special_functions/stirling.lean | 4 +++- src/data/real/cardinality.lean | 10 +++++++--- src/data/real/irrational.lean | 2 +- src/number_theory/padics/padic_integers.lean | 2 +- src/topology/metric_space/lipschitz.lean | 4 ++-- src/topology/metric_space/pi_nat.lean | 1 + 9 files changed, 22 insertions(+), 14 deletions(-) diff --git a/src/analysis/special_functions/bernstein.lean b/src/analysis/special_functions/bernstein.lean index e84fbc01ca3d0..67df5d5c9e191 100644 --- a/src/analysis/special_functions/bernstein.lean +++ b/src/analysis/special_functions/bernstein.lean @@ -192,7 +192,8 @@ lemma le_of_mem_S_compl (1 : ℝ) ≤ (δ f ε h)^(-2 : ℤ) * (x - k/ₙ) ^ 2 := begin simp only [finset.mem_compl, not_lt, set.mem_to_finset, set.mem_set_of_eq, S] at m, - erw [zpow_neg, ← div_eq_inv_mul, one_le_div (pow_pos δ_pos 2), sq_le_sq, abs_of_pos δ_pos], + rw [zpow_neg, ← div_eq_inv_mul, zpow_two, ←pow_two, one_le_div (pow_pos δ_pos 2), sq_le_sq, + abs_of_pos δ_pos], rwa [dist_comm] at m end diff --git a/src/analysis/special_functions/integrals.lean b/src/analysis/special_functions/integrals.lean index 444560c44e200..efd34ff0461cc 100644 --- a/src/analysis/special_functions/integrals.lean +++ b/src/analysis/special_functions/integrals.lean @@ -265,7 +265,7 @@ begin end @[simp] lemma integral_pow : ∫ x in a..b, x ^ n = (b ^ (n + 1) - a ^ (n + 1)) / (n + 1) := -by simpa using integral_zpow (or.inl (int.coe_nat_nonneg n)) +by simpa only [←int.coe_nat_succ, zpow_coe_nat] using integral_zpow (or.inl (int.coe_nat_nonneg n)) /-- Integral of `|x - a| ^ n` over `Ι a b`. This integral appears in the proof of the Picard-Lindelöf/Cauchy-Lipschitz theorem. -/ diff --git a/src/analysis/special_functions/pow.lean b/src/analysis/special_functions/pow.lean index 4f46b791aee68..4fd66dc7d3314 100644 --- a/src/analysis/special_functions/pow.lean +++ b/src/analysis/special_functions/pow.lean @@ -570,13 +570,13 @@ by rw [rpow_def, complex.of_real_add, complex.cpow_add _ _ (complex.of_real_ne_z complex.of_real_mul_re, ← rpow_def, mul_comm] lemma rpow_add_nat {x : ℝ} (hx : x ≠ 0) (y : ℝ) (n : ℕ) : x ^ (y + n) = x ^ y * x ^ n := -rpow_add_int hx y n +by simpa using rpow_add_int hx y n lemma rpow_sub_int {x : ℝ} (hx : x ≠ 0) (y : ℝ) (n : ℤ) : x ^ (y - n) = x ^ y / x ^ n := by simpa using rpow_add_int hx y (-n) lemma rpow_sub_nat {x : ℝ} (hx : x ≠ 0) (y : ℝ) (n : ℕ) : x ^ (y - n) = x ^ y / x ^ n := -rpow_sub_int hx y n +by simpa using rpow_sub_int hx y n lemma rpow_add_one {x : ℝ} (hx : x ≠ 0) (y : ℝ) : x ^ (y + 1) = x ^ y * x := by simpa using rpow_add_nat hx y 1 @@ -589,7 +589,7 @@ by simp only [rpow_def, ← complex.of_real_zpow, complex.cpow_int_cast, complex.of_real_int_cast, complex.of_real_re] @[simp, norm_cast] lemma rpow_nat_cast (x : ℝ) (n : ℕ) : x ^ (n : ℝ) = x ^ n := -rpow_int_cast x n +by simpa using rpow_int_cast x n @[simp] lemma rpow_two (x : ℝ) : x ^ (2 : ℝ) = x ^ 2 := by { rw ← rpow_nat_cast, simp only [nat.cast_bit0, nat.cast_one] } @@ -1121,7 +1121,7 @@ by simpa only [rpow_int_cast] using is_o_rpow_exp_pos_mul_at_top k hb /-- `x ^ k = o(exp(b * x))` as `x → ∞` for any natural `k` and positive `b`. -/ lemma is_o_pow_exp_pos_mul_at_top (k : ℕ) {b : ℝ} (hb : 0 < b) : (λ x : ℝ, x ^ k) =o[at_top] (λ x, exp (b * x)) := -is_o_zpow_exp_pos_mul_at_top k hb +by simpa using is_o_zpow_exp_pos_mul_at_top k hb /-- `x ^ s = o(exp x)` as `x → ∞` for any real `s`. -/ lemma is_o_rpow_exp_at_top (s : ℝ) : (λ x : ℝ, x ^ s) =o[at_top] exp := diff --git a/src/analysis/special_functions/stirling.lean b/src/analysis/special_functions/stirling.lean index ae5e25a3b556b..01a297e311410 100644 --- a/src/analysis/special_functions/stirling.lean +++ b/src/analysis/special_functions/stirling.lean @@ -104,7 +104,9 @@ begin have h_nonneg : 0 ≤ ((1 / (2 * (n.succ : ℝ) + 1)) ^ 2) := sq_nonneg _, have g : has_sum (λ k : ℕ, ((1 / (2 * (n.succ : ℝ) + 1)) ^ 2) ^ k.succ) ((1 / (2 * n.succ + 1)) ^ 2 / (1 - (1 / (2 * n.succ + 1)) ^ 2)), - { refine (has_sum_geometric_of_lt_1 h_nonneg _).mul_left ((1 / (2 * (n.succ : ℝ) + 1)) ^ 2), + { have := (has_sum_geometric_of_lt_1 h_nonneg _).mul_left ((1 / (2 * (n.succ : ℝ) + 1)) ^ 2), + { simp_rw ←pow_succ at this, + exact this, }, rw [one_div, inv_pow], exact inv_lt_one (one_lt_pow ((lt_add_iff_pos_left 1).mpr $ by positivity) two_ne_zero) }, have hab : ∀ (k : ℕ), (1 / (2 * (k.succ : ℝ) + 1)) * ((1 / (2 * n.succ + 1)) ^ 2) ^ k.succ ≤ diff --git a/src/data/real/cardinality.lean b/src/data/real/cardinality.lean index 58db59e8ab65a..74c5933f30a7f 100644 --- a/src/data/real/cardinality.lean +++ b/src/data/real/cardinality.lean @@ -66,6 +66,10 @@ lemma cantor_function_aux_eq (h : f n = g n) : cantor_function_aux c f n = cantor_function_aux c g n := by simp [cantor_function_aux, h] +lemma cantor_function_aux_zero (f : ℕ → bool) : + cantor_function_aux c f 0 = cond (f 0) 1 0 := +by { cases h : f 0; simp [h] } + lemma cantor_function_aux_succ (f : ℕ → bool) : (λ n, cantor_function_aux c f (n + 1)) = λ n, c * cantor_function_aux c (λ n, f (n + 1)) n := by { ext n, cases h : f (n + 1); simp [h, pow_succ] } @@ -122,9 +126,9 @@ begin { rw [cantor_function_succ _ (le_of_lt h1) h3, div_eq_mul_inv, ←tsum_geometric_of_lt_1 (le_of_lt h1) h3], apply zero_add }, - { convert tsum_eq_single 0 _, - { apply_instance }, - { intros n hn, cases n, contradiction, refl } } }, + { refine (tsum_eq_single 0 _).trans _, + { intros n hn, cases n, contradiction, refl }, + { exact cantor_function_aux_zero _ }, } }, rw [cantor_function_succ f (le_of_lt h1) h3, cantor_function_succ g (le_of_lt h1) h3], rw [hn 0 $ zero_lt_succ n], apply add_lt_add_left, rw mul_lt_mul_left h1, exact ih (λ k hk, hn _ $ nat.succ_lt_succ hk) fn gn diff --git a/src/data/real/irrational.lean b/src/data/real/irrational.lean index 6c67f6fbc36ef..c2463026fe019 100644 --- a/src/data/real/irrational.lean +++ b/src/data/real/irrational.lean @@ -351,7 +351,7 @@ theorem of_pow : ∀ n : ℕ, irrational (x^n) → irrational x | (n+1) := λ h, by { rw pow_succ at h, exact h.mul_cases.elim id (of_pow n) } theorem of_zpow : ∀ m : ℤ, irrational (x^m) → irrational x -| (n:ℕ) := of_pow n +| (n:ℕ) := λ h, by { rw zpow_coe_nat at h, exact h.of_pow _ } | -[1+n] := λ h, by { rw zpow_neg_succ_of_nat at h, exact h.of_inv.of_pow _ } end irrational diff --git a/src/number_theory/padics/padic_integers.lean b/src/number_theory/padics/padic_integers.lean index 5b025a364f284..b70cfcbe73956 100644 --- a/src/number_theory/padics/padic_integers.lean +++ b/src/number_theory/padics/padic_integers.lean @@ -314,7 +314,7 @@ begin { simp [hx] }, have h : (1 : ℝ) < p := by exact_mod_cast hp.1.one_lt, rw [← neg_nonpos, ← (zpow_strict_mono h).le_iff_le], - show (p : ℝ) ^ -valuation x ≤ p ^ 0, + show (p : ℝ) ^ -valuation x ≤ p ^ (0 : ℤ), rw [← norm_eq_pow_val hx], simpa using x.property end diff --git a/src/topology/metric_space/lipschitz.lean b/src/topology/metric_space/lipschitz.lean index 724a89cbada33..4f33443074d1e 100644 --- a/src/topology/metric_space/lipschitz.lean +++ b/src/topology/metric_space/lipschitz.lean @@ -240,7 +240,7 @@ end protected lemma iterate {f : α → α} (hf : lipschitz_with K f) : ∀n, lipschitz_with (K ^ n) (f^[n]) -| 0 := lipschitz_with.id +| 0 := by simpa only [pow_zero] using lipschitz_with.id | (n + 1) := by rw [pow_succ']; exact (iterate n).comp hf lemma edist_iterate_succ_le_geometric {f : α → α} (hf : lipschitz_with K f) (x n) : @@ -265,7 +265,7 @@ protected lemma list_prod (f : ι → function.End α) (K : ι → ℝ≥0) protected lemma pow {f : function.End α} {K} (h : lipschitz_with K f) : ∀ n : ℕ, lipschitz_with (K^n) (f^n : function.End α) -| 0 := lipschitz_with.id +| 0 := by simpa only [pow_zero] using lipschitz_with.id | (n + 1) := by { rw [pow_succ, pow_succ], exact h.mul (pow n) } end emetric diff --git a/src/topology/metric_space/pi_nat.lean b/src/topology/metric_space/pi_nat.lean index 0a7287c0e7b4f..ba35f8bb48a1a 100644 --- a/src/topology/metric_space/pi_nat.lean +++ b/src/topology/metric_space/pi_nat.lean @@ -709,6 +709,7 @@ begin apply apply_first_diff_ne hne', rw [le_zero_iff.1 h], apply apply_eq_of_dist_lt _ le_rfl, + rw pow_zero, exact hxy }, have hn : first_diff x.1 y.1 = n + 1 := (nat.succ_pred_eq_of_pos diff_pos).symm, rw [dist', dist_eq_of_ne hne', hn], From d50b12ae8e2bd910d08a94823976adae9825718b Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 12 Jan 2023 20:46:51 +0000 Subject: [PATCH 0225/1291] fix(*): fix typo in the word subtraction (#18145) Wiktionary defines "substraction" as "(obsolete or non-native speakers' English) Subtraction." --- archive/oxford_invariants/2021summer/week3_p1.lean | 4 ++-- src/analysis/normed/field/basic.lean | 4 ++-- src/data/nat/dist.lean | 2 +- src/data/nat/factorial/cast.lean | 6 +++--- src/data/num/bitwise.lean | 2 +- src/topology/algebra/infinite_sum.lean | 4 ++-- 6 files changed, 11 insertions(+), 11 deletions(-) diff --git a/archive/oxford_invariants/2021summer/week3_p1.lean b/archive/oxford_invariants/2021summer/week3_p1.lean index d9958cd21cac4..d0a8fb0731ff8 100644 --- a/archive/oxford_invariants/2021summer/week3_p1.lean +++ b/archive/oxford_invariants/2021summer/week3_p1.lean @@ -22,7 +22,7 @@ Mathlib is based on type theory, so saying that a rational is a natural doesn't we ask that there exists `b : ℕ` whose cast to `α` is the sum we want. In mathlib, `ℕ` starts at `0`. To make the indexing cleaner, we use `a₀, ..., aₙ₋₁` instead of -`a₁, ..., aₙ`. Similarly, it's nicer to not use substraction of naturals, so we replace +`a₁, ..., aₙ`. Similarly, it's nicer to not use subtraction of naturals, so we replace `aᵢ ∣ aᵢ₋₁ + aᵢ₊₁` by `aᵢ₊₁ ∣ aᵢ + aᵢ₊₂`. We don't actually have to work in `ℚ` or `ℝ`. We can be even more general by stating the result for @@ -106,7 +106,7 @@ begin obtain ⟨b, hb, han⟩ := ih (λ i hi, ha i $ nat.le_succ_of_le hi) (λ i hi, a_pos i $ nat.le_succ_of_le hi), specialize ha n le_rfl, - have ha₀ : a 0 ≤ a n * b, -- Needing this is an artifact of `ℕ`-substraction. + have ha₀ : a 0 ≤ a n * b, -- Needing this is an artifact of `ℕ`-subtraction. { rw [←@nat.cast_le α, nat.cast_mul, hb, ←div_le_iff' (a_pos _ $ n.le_succ.trans $ nat.le_succ _), ←mul_div_mul_right _ _ (a_pos _ $ nat.le_succ _).ne'], suffices h : ∀ i, i ∈ finset.range (n + 1) → 0 ≤ (a 0 : α) * a (n + 1) / (a i * a (i + 1)), diff --git a/src/analysis/normed/field/basic.lean b/src/analysis/normed/field/basic.lean index d7dd051897b82..f610971286317 100644 --- a/src/analysis/normed/field/basic.lean +++ b/src/analysis/normed/field/basic.lean @@ -805,8 +805,8 @@ tsum_mul_tsum (summable_of_summable_norm hf) (summable_of_summable_norm hg) We prove two versions of the Cauchy product formula. The first one is `tsum_mul_tsum_eq_tsum_sum_range_of_summable_norm`, where the `n`-th term is a sum over -`finset.range (n+1)` involving `nat` substraction. -In order to avoid `nat` substraction, we also provide +`finset.range (n+1)` involving `nat` subtraction. +In order to avoid `nat` subtraction, we also provide `tsum_mul_tsum_eq_tsum_sum_antidiagonal_of_summable_norm`, where the `n`-th term is a sum over all pairs `(k, l)` such that `k+l=n`, which corresponds to the `finset` `finset.nat.antidiagonal n`. -/ diff --git a/src/data/nat/dist.lean b/src/data/nat/dist.lean index 5b32346de16e6..318fa77a2ba14 100644 --- a/src/data/nat/dist.lean +++ b/src/data/nat/dist.lean @@ -11,7 +11,7 @@ import data.nat.order.basic > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. > Any changes to this file require a corresponding PR to mathlib4. -This file defines a simple distance function on naturals from truncated substraction. +This file defines a simple distance function on naturals from truncated subtraction. -/ namespace nat diff --git a/src/data/nat/factorial/cast.lean b/src/data/nat/factorial/cast.lean index 27a4a241076ad..9acc2393ce578 100644 --- a/src/data/nat/factorial/cast.lean +++ b/src/data/nat/factorial/cast.lean @@ -11,8 +11,8 @@ import ring_theory.polynomial.pochhammer This file allows calculating factorials (including ascending and descending ones) as elements of a semiring. -This is particularly crucial for `nat.desc_factorial` as substraction on `ℕ` does **not** correspond -to substraction on a general semiring. For example, we can't rely on existing cast lemmas to prove +This is particularly crucial for `nat.desc_factorial` as subtraction on `ℕ` does **not** correspond +to subtraction on a general semiring. For example, we can't rely on existing cast lemmas to prove `↑(a.desc_factorial 2) = ↑a * (↑a - 1)`. We must use the fact that, whenever `↑(a - 1)` is not equal to `↑a - 1`, the other factor is `0` anyway. -/ @@ -52,7 +52,7 @@ end semiring section ring variables [ring S] (a b : ℕ) -/-- Convenience lemma. The `a - 1` is not using truncated substraction, as opposed to the definition +/-- Convenience lemma. The `a - 1` is not using truncated subtraction, as opposed to the definition of `nat.desc_factorial` as a natural. -/ lemma cast_desc_factorial_two : (a.desc_factorial 2 : S) = a * (a - 1) := diff --git a/src/data/num/bitwise.lean b/src/data/num/bitwise.lean index 63b576b02521a..d88123716468b 100644 --- a/src/data/num/bitwise.lean +++ b/src/data/num/bitwise.lean @@ -339,7 +339,7 @@ protected def add (a b : snum) : snum := cadd a b ff instance : has_add snum := ⟨snum.add⟩ -/-- Substract two `snum`s. -/ +/-- subtract two `snum`s. -/ protected def sub (a b : snum) : snum := a + -b instance : has_sub snum := ⟨snum.sub⟩ diff --git a/src/topology/algebra/infinite_sum.lean b/src/topology/algebra/infinite_sum.lean index 8e9bccace25ac..885d26f843fd5 100644 --- a/src/topology/algebra/infinite_sum.lean +++ b/src/topology/algebra/infinite_sum.lean @@ -1612,8 +1612,8 @@ section cauchy_product We prove two versions of the Cauchy product formula. The first one is `tsum_mul_tsum_eq_tsum_sum_range`, where the `n`-th term is a sum over `finset.range (n+1)` -involving `nat` substraction. -In order to avoid `nat` substraction, we also provide `tsum_mul_tsum_eq_tsum_sum_antidiagonal`, +involving `nat` subtraction. +In order to avoid `nat` subtraction, we also provide `tsum_mul_tsum_eq_tsum_sum_antidiagonal`, where the `n`-th term is a sum over all pairs `(k, l)` such that `k+l=n`, which corresponds to the `finset` `finset.nat.antidiagonal n` -/ From feb165c980c918bb296fede8c3b21dbb4b85bb56 Mon Sep 17 00:00:00 2001 From: Devon Tuma Date: Fri, 13 Jan 2023 00:08:11 +0000 Subject: [PATCH 0226/1291] chore(probability_mass_function/monad): Use standard `Union` notation (#18154) Use `Union` instead of `set_of` in `support` of `bind` operations. --- src/probability/probability_mass_function/monad.lean | 10 +++++----- 1 file changed, 5 insertions(+), 5 deletions(-) diff --git a/src/probability/probability_mass_function/monad.lean b/src/probability/probability_mass_function/monad.lean index 49b3980aab14c..76cf6758ed6f0 100644 --- a/src/probability/probability_mass_function/monad.lean +++ b/src/probability/probability_mass_function/monad.lean @@ -74,11 +74,11 @@ variables (p : pmf α) (f : α → pmf β) (g : β → pmf γ) @[simp] lemma bind_apply (b : β) : p.bind f b = ∑'a, p a * f a b := rfl -@[simp] lemma support_bind : (p.bind f).support = {b | ∃ a ∈ p.support, b ∈ (f a).support} := +@[simp] lemma support_bind : (p.bind f).support = ⋃ a ∈ p.support, (f a).support := set.ext (λ b, by simp [mem_support_iff, ennreal.tsum_eq_zero, not_or_distrib]) lemma mem_support_bind_iff (b : β) : b ∈ (p.bind f).support ↔ ∃ a ∈ p.support, b ∈ (f a).support := -by simp only [support_bind, set.mem_set_of_eq] +by simp only [support_bind, set.mem_Union, set.mem_set_of_eq] @[simp] lemma pure_bind (a : α) (f : α → pmf β) : (pure a).bind f = f a := have ∀ b a', ite (a' = a) 1 0 * f a' b = ite (a' = a) (f a b) 0, from @@ -155,18 +155,18 @@ variables {p : pmf α} (f : Π a ∈ p.support, pmf β) p.bind_on_support f b = ∑' a, p a * if h : p a = 0 then 0 else f a h b := rfl @[simp] lemma support_bind_on_support : - (p.bind_on_support f).support = {b | ∃ (a : α) (h : a ∈ p.support), b ∈ (f a h).support} := + (p.bind_on_support f).support = ⋃ (a : α) (h : a ∈ p.support), (f a h).support := begin refine set.ext (λ b, _), simp only [ennreal.tsum_eq_zero, not_or_distrib, mem_support_iff, - bind_on_support_apply, ne.def, not_forall, mul_eq_zero], + bind_on_support_apply, ne.def, not_forall, mul_eq_zero, set.mem_Union], exact ⟨λ hb, let ⟨a, ⟨ha, ha'⟩⟩ := hb in ⟨a, ha, by simpa [ha] using ha'⟩, λ hb, let ⟨a, ha, ha'⟩ := hb in ⟨a, ⟨ha, by simpa [(mem_support_iff _ a).1 ha] using ha'⟩⟩⟩ end lemma mem_support_bind_on_support_iff (b : β) : b ∈ (p.bind_on_support f).support ↔ ∃ (a : α) (h : a ∈ p.support), b ∈ (f a h).support := -by rw [support_bind_on_support, set.mem_set_of_eq] +by simp only [support_bind_on_support, set.mem_set_of_eq, set.mem_Union] /-- `bind_on_support` reduces to `bind` if `f` doesn't depend on the additional hypothesis -/ @[simp] lemma bind_on_support_eq_bind (p : pmf α) (f : α → pmf β) : From f3187269ad18e82a809428a42d6282ce81e4ebcc Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Fri, 13 Jan 2023 02:46:46 +0000 Subject: [PATCH 0227/1291] feat(data/multiset/basic): `multiset.coe_eq_singleton` (#15910) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Yaël Dillies --- src/data/multiset/basic.lean | 11 ++++++++++- 1 file changed, 10 insertions(+), 1 deletion(-) diff --git a/src/data/multiset/basic.lean b/src/data/multiset/basic.lean index ed38dc1edbfe5..d6f6ab88f8906 100644 --- a/src/data/multiset/basic.lean +++ b/src/data/multiset/basic.lean @@ -227,7 +227,7 @@ instance : is_lawful_singleton α (multiset α) := ⟨λ a, rfl⟩ @[simp] theorem cons_zero (a : α) : a ::ₘ 0 = {a} := rfl -@[simp] theorem coe_singleton (a : α) : ([a] : multiset α) = {a} := rfl +@[simp, norm_cast] theorem coe_singleton (a : α) : ([a] : multiset α) = {a} := rfl @[simp] theorem mem_singleton {a b : α} : b ∈ ({a} : multiset α) ↔ b = a := by simp only [←cons_zero, mem_cons, iff_self, or_false, not_mem_zero] @@ -238,6 +238,9 @@ by { rw ←cons_zero, exact mem_cons_self _ _ } @[simp] theorem singleton_inj {a b : α} : ({a} : multiset α) = {b} ↔ a = b := by { simp_rw [←cons_zero], exact cons_inj_left _ } +@[simp, norm_cast] lemma coe_eq_singleton {l : list α} {a : α} : (l : multiset α) = {a} ↔ l = [a] := +by rw [←coe_singleton, coe_eq_coe, list.perm_singleton] + @[simp] lemma singleton_eq_cons_iff {a b : α} (m : multiset α) : {a} = b ::ₘ m ↔ a = b ∧ m = 0 := by { rw [←cons_zero, cons_eq_cons], simp [eq_comm] } @@ -315,6 +318,12 @@ empty_iff_eq_nil.trans to_list_eq_nil @[simp] lemma mem_to_list {a : α} {s : multiset α} : a ∈ s.to_list ↔ a ∈ s := by rw [← mem_coe, coe_to_list] +@[simp] lemma to_list_eq_singleton_iff {a : α} {m : multiset α} : m.to_list = [a] ↔ m = {a} := +by rw [←perm_singleton, ←coe_eq_coe, coe_to_list, coe_singleton] + +@[simp] lemma to_list_singleton (a : α) : ({a} : multiset α).to_list = [a] := +multiset.to_list_eq_singleton_iff.2 rfl + end to_list /-! ### Partial order on `multiset`s -/ From 527e406f1729eed5f9960eb25b68b73115fdd9fb Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Fri, 13 Jan 2023 05:43:49 +0000 Subject: [PATCH 0228/1291] chore(data/list): golf, merge 2 files (#18120) - merge `data.list.modeq` into `data.list.rotate`; - mark `list.rotate_eq_rotate` as `@[simp]`; - golf some proofs. --- src/data/list/modeq.lean | 68 ------------------------- src/data/list/rotate.lean | 73 +++++++++++++++++---------- src/group_theory/perm/cycle/type.lean | 2 +- 3 files changed, 47 insertions(+), 96 deletions(-) delete mode 100644 src/data/list/modeq.lean diff --git a/src/data/list/modeq.lean b/src/data/list/modeq.lean deleted file mode 100644 index 10121abd06831..0000000000000 --- a/src/data/list/modeq.lean +++ /dev/null @@ -1,68 +0,0 @@ -/- -Copyright (c) 2018 Chris Hughes. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Chris Hughes --/ -import data.nat.modeq -import data.list.rotate - -/-! # List rotation and modular arithmetic -/ - -namespace list -variable {α : Type*} - -lemma nth_rotate : ∀ {l : list α} {n m : ℕ} (hml : m < l.length), - (l.rotate n).nth m = l.nth ((m + n) % l.length) -| [] n m hml := (nat.not_lt_zero _ hml).elim -| l 0 m hml := by simp [nat.mod_eq_of_lt hml] -| (a::l) (n+1) m hml := -have h₃ : m < list.length (l ++ [a]), by simpa using hml, -(lt_or_eq_of_le (nat.le_of_lt_succ $ nat.mod_lt (m + n) - (lt_of_le_of_lt (nat.zero_le _) hml))).elim -(λ hml', - have h₁ : (m + (n + 1)) % ((a :: l : list α).length) = - (m + n) % ((a :: l : list α).length) + 1, - from calc (m + (n + 1)) % (l.length + 1) = - ((m + n) % (l.length + 1) + 1) % (l.length + 1) : - add_assoc m n 1 ▸ nat.modeq.add_right 1 (nat.mod_mod _ _).symm - ... = (m + n) % (l.length + 1) + 1 : nat.mod_eq_of_lt (nat.succ_lt_succ hml'), - have h₂ : (m + n) % (l ++ [a]).length < l.length, by simpa [nat.add_one] using hml', - by rw [list.rotate_cons_succ, nth_rotate h₃, list.nth_append h₂, h₁, list.nth]; simp) -(λ hml', - have h₁ : (m + (n + 1)) % (l.length + 1) = 0, - from calc (m + (n + 1)) % (l.length + 1) = (l.length + 1) % (l.length + 1) : - add_assoc m n 1 ▸ nat.modeq.add_right 1 - (hml'.trans (nat.mod_eq_of_lt (nat.lt_succ_self _)).symm) - ... = 0 : by simp, - by rw [list.length, list.rotate_cons_succ, nth_rotate h₃, list.length_append, - list.length_cons, list.length, zero_add, hml', h₁, list.nth_concat_length]; refl) - -lemma rotate_eq_self_iff_eq_repeat [hα : nonempty α] : ∀ {l : list α}, - (∀ n, l.rotate n = l) ↔ ∃ a, l = list.repeat a l.length -| [] := ⟨λ h, nonempty.elim hα (λ a, ⟨a, by simp⟩), by simp⟩ -| (a::l) := -⟨λ h, ⟨a, list.ext_le (by simp) $ λ n hn h₁, - begin - rw [← option.some_inj, ← list.nth_le_nth], - conv {to_lhs, rw ← h ((list.length (a :: l)) - n)}, - rw [nth_rotate hn, add_tsub_cancel_of_le (le_of_lt hn), - nat.mod_self, nth_le_repeat], refl - end⟩, - λ ⟨a, ha⟩ n, ha.symm ▸ list.ext_le (by simp) - (λ m hm h, - have hm' : (m + n) % (list.repeat a (list.length (a :: l))).length < list.length (a :: l), - by rw list.length_repeat; exact nat.mod_lt _ (nat.succ_pos _), - by rw [nth_le_repeat, ← option.some_inj, ← list.nth_le_nth, nth_rotate h, list.nth_le_nth, - nth_le_repeat]; simp * at *)⟩ - -lemma rotate_repeat (a : α) (n : ℕ) (k : ℕ) : - (list.repeat a n).rotate k = list.repeat a n := -let h : nonempty α := ⟨a⟩ in by exactI rotate_eq_self_iff_eq_repeat.mpr ⟨a, by rw length_repeat⟩ k - -lemma rotate_one_eq_self_iff_eq_repeat [nonempty α] {l : list α} : - l.rotate 1 = l ↔ ∃ a : α, l = list.repeat a l.length := -⟨λ h, rotate_eq_self_iff_eq_repeat.mp (λ n, nat.rec l.rotate_zero - (λ n hn, by rwa [nat.succ_eq_add_one, ←l.rotate_rotate, hn]) n), - λ h, rotate_eq_self_iff_eq_repeat.mpr h 1⟩ - -end list diff --git a/src/data/list/rotate.lean b/src/data/list/rotate.lean index b1ba189819b28..17fbef984564c 100644 --- a/src/data/list/rotate.lean +++ b/src/data/list/rotate.lean @@ -24,7 +24,7 @@ rotated, rotation, permutation, cycle universe u variables {α : Type u} -open nat +open nat function namespace list @@ -99,6 +99,11 @@ by rw [rotate_eq_rotate', rotate_eq_rotate', rotate'_cons_succ] @[simp] lemma length_rotate (l : list α) (n : ℕ) : (l.rotate n).length = l.length := by rw [rotate_eq_rotate', length_rotate'] +lemma rotate_repeat (a : α) (n : ℕ) (k : ℕ) : + (repeat a n).rotate k = repeat a n := +eq_repeat.2 ⟨by rw [length_rotate, length_repeat], + λ b hb, eq_of_mem_repeat $ mem_rotate.1 hb⟩ + lemma rotate_eq_drop_append_take {l : list α} {n : ℕ} : n ≤ l.length → l.rotate n = l.drop n ++ l.take n := by rw rotate_eq_rotate'; exact rotate'_eq_drop_append_take @@ -165,23 +170,7 @@ by rw [eq_comm, rotate_eq_nil_iff, eq_comm] @[simp] lemma rotate_singleton (x : α) (n : ℕ) : [x].rotate n = [x] := -begin - induction n with n hn, - { simp }, - { rwa [rotate_cons_succ] } -end - -@[simp] lemma rotate_eq_singleton_iff {l : list α} {n : ℕ} {x : α} : l.rotate n = [x] ↔ l = [x] := -begin - induction n with n hn generalizing l, - { simp }, - { cases l with hd tl, - { simp }, - { simp [rotate_cons_succ, hn, append_eq_cons_iff, and_comm] } } -end - -@[simp] lemma singleton_eq_rotate_iff {l : list α} {n : ℕ} {x : α} : [x] = l.rotate n ↔ [x] = l := -by rw [eq_comm, rotate_eq_singleton_iff, eq_comm] +rotate_repeat x 1 n lemma zip_with_rotate_distrib {α β γ : Type*} (f : α → β → γ) (l : list α) (l' : list β) (n : ℕ) (h : l.length = l'.length) : @@ -241,18 +230,42 @@ begin simpa [mod_eq_of_lt hm, tsub_add_cancel_of_le hn'.le] using nat.mod_eq_of_lt hk } end -lemma rotate_injective (n : ℕ) : function.injective (λ l : list α, l.rotate n) := +lemma nth_rotate {l : list α} {n m : ℕ} (hml : m < l.length) : + (l.rotate n).nth m = l.nth ((m + n) % l.length) := begin - rintros l l' (h : l.rotate n = l'.rotate n), - have hle : l.length = l'.length := (l.length_rotate n).symm.trans (h.symm ▸ l'.length_rotate n), - rw [rotate_eq_drop_append_take_mod, rotate_eq_drop_append_take_mod] at h, - obtain ⟨hd, ht⟩ := append_inj h _, - { rw [←take_append_drop _ l, ht, hd, take_append_drop] }, - { rw [length_drop, length_drop, hle] } + rw [nth_le_nth, nth_le_nth (nat.mod_lt _ _), nth_le_rotate], + rwa [length_rotate] end --- possibly easier to find in doc-gen, otherwise not that useful. -lemma rotate_eq_rotate {l l' : list α} {n : ℕ} : +lemma head'_rotate {l : list α} {n : ℕ} (h : n < l.length) : + head' (l.rotate n) = l.nth n := +by rw [← nth_zero, nth_rotate (n.zero_le.trans_lt h), zero_add, nat.mod_eq_of_lt h] + +lemma rotate_eq_self_iff_eq_repeat [hα : nonempty α] : + ∀ {l : list α}, (∀ n, l.rotate n = l) ↔ ∃ a, l = repeat a l.length +| [] := by simp +| (a :: l) := ⟨λ h, ⟨a, ext_le (length_repeat _ _).symm $ λ n h₁ h₂, + begin + inhabit α, + rw [nth_le_repeat, ← option.some_inj, ← nth_le_nth, ← head'_rotate h₁, h, head'] + end⟩, λ ⟨b, hb⟩ n, by rw [hb, rotate_repeat]⟩ + +lemma rotate_one_eq_self_iff_eq_repeat [nonempty α] {l : list α} : + l.rotate 1 = l ↔ ∃ a : α, l = repeat a l.length := +⟨λ h, rotate_eq_self_iff_eq_repeat.mp (λ n, nat.rec l.rotate_zero + (λ n hn, by rwa [nat.succ_eq_add_one, ←l.rotate_rotate, hn]) n), + λ h, rotate_eq_self_iff_eq_repeat.mpr h 1⟩ + +lemma rotate_injective (n : ℕ) : injective (λ l : list α, l.rotate n) := +begin + rintro l₁ l₂ (h : l₁.rotate n = l₂.rotate n), + have : length l₁ = length l₂, by simpa only [length_rotate] using congr_arg length h, + refine ext_le this (λ k h₁ h₂, _), + rw [← nth_le_rotate' l₁ n, ← nth_le_rotate' l₂ n], + congr' 1; simp only [h, this] +end + +@[simp] lemma rotate_eq_rotate {l l' : list α} {n : ℕ} : l.rotate n = l'.rotate n ↔ l = l' := (rotate_injective n).eq_iff @@ -268,6 +281,12 @@ begin exact (nat.mod_lt _ hl).le } } end +@[simp] lemma rotate_eq_singleton_iff {l : list α} {n : ℕ} {x : α} : l.rotate n = [x] ↔ l = [x] := +⟨λ h, by rw [rotate_eq_iff.1 h, rotate_singleton], λ h, h.symm ▸ rotate_singleton _ _⟩ + +@[simp] lemma singleton_eq_rotate_iff {l : list α} {n : ℕ} {x : α} : [x] = l.rotate n ↔ [x] = l := +by rw [eq_comm, rotate_eq_singleton_iff, eq_comm] + lemma reverse_rotate (l : list α) (n : ℕ) : (l.rotate n).reverse = l.reverse.rotate (l.length - (n % l.length)) := begin diff --git a/src/group_theory/perm/cycle/type.lean b/src/group_theory/perm/cycle/type.lean index 9069d0086b584..1aac2fe826279 100644 --- a/src/group_theory/perm/cycle/type.lean +++ b/src/group_theory/perm/cycle/type.lean @@ -6,7 +6,7 @@ Authors: Thomas Browning import algebra.gcd_monoid.multiset import combinatorics.partition -import data.list.modeq +import data.list.rotate import group_theory.perm.cycle.basic import ring_theory.int.basic import tactic.linarith From f6a7bd6f0540682385e63620c65e03cb076695ea Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Fri, 13 Jan 2023 05:43:52 +0000 Subject: [PATCH 0229/1291] feat(data/finset/basic): add a `can_lift` instance (#18152) --- src/data/finset/basic.lean | 4 ++++ 1 file changed, 4 insertions(+) diff --git a/src/data/finset/basic.lean b/src/data/finset/basic.lean index cb20eea268afe..0966010e09b20 100644 --- a/src/data/finset/basic.lean +++ b/src/data/finset/basic.lean @@ -138,6 +138,10 @@ structure finset (α : Type*) := (val : multiset α) (nodup : nodup val) +instance multiset.can_lift_finset {α} : + can_lift (multiset α) (finset α) finset.val multiset.nodup := +⟨λ m hm, ⟨⟨m, hm⟩, rfl⟩⟩ + namespace finset theorem eq_of_veq : ∀ {s t : finset α}, s.1 = t.1 → s = t From d7943885503965d07ccaeb390d65fbc3f45962e6 Mon Sep 17 00:00:00 2001 From: leanprover-community-bot Date: Fri, 13 Jan 2023 05:43:53 +0000 Subject: [PATCH 0230/1291] chore(scripts): update nolints.txt (#18161) I am happy to remove some nolints for you! --- scripts/nolints.txt | 4 +++- 1 file changed, 3 insertions(+), 1 deletion(-) diff --git a/scripts/nolints.txt b/scripts/nolints.txt index 2b9acf01a96c8..60d36e81c535d 100644 --- a/scripts/nolints.txt +++ b/scripts/nolints.txt @@ -352,9 +352,11 @@ apply_nolint affine_independent.exists_unique_dist_eq fintype_finite apply_nolint sub_mul_action.has_zero fails_quickly -- linear_algebra/affine_space/basis.lean +apply_nolint affine_basis.ext_elem fintype_finite + +-- linear_algebra/affine_space/matrix.lean apply_nolint affine_basis.affine_independent_of_to_matrix_right_inv fintype_finite apply_nolint affine_basis.affine_span_eq_top_of_to_matrix_left_inv fintype_finite -apply_nolint affine_basis.ext_elem fintype_finite -- logic/relator.lean apply_nolint relator.bi_total doc_blame From cf8e77c636317b059a8ce20807a29cf3772a0640 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 13 Jan 2023 10:23:39 +0000 Subject: [PATCH 0231/1291] feat(group_theory/subgroup/basic): Map a subgroup under an iso (#17775) Cross-reference `mul_equiv.subgroup_map` and `subgroup.equiv_map_of_injective`. Add lemmas connecting them. --- src/group_theory/commensurable.lean | 6 +-- src/group_theory/subgroup/basic.lean | 48 +++++++++++++--------- src/group_theory/submonoid/operations.lean | 26 ++++++++---- 3 files changed, 51 insertions(+), 29 deletions(-) diff --git a/src/group_theory/commensurable.lean b/src/group_theory/commensurable.lean index 73e3da604faba..1d7772b62ef73 100644 --- a/src/group_theory/commensurable.lean +++ b/src/group_theory/commensurable.lean @@ -52,10 +52,10 @@ lemma equivalence : equivalence (@commensurable G _) := /--Equivalence of `K/H ⊓ K` with `gKg⁻¹/gHg⁻¹ ⊓ gKg⁻¹`-/ def quot_conj_equiv (H K : subgroup G) (g : conj_act G) : K ⧸ (H.subgroup_of K) ≃ (g • K).1 ⧸ ((g • H).subgroup_of (g • K)) := -quotient.congr (K.equiv_smul g).to_equiv (λ a b, by { rw [←quotient.eq', ←quotient.eq', +quotient.congr (K.equiv_smul g).to_equiv (λ a b, by { dsimp, rw [←quotient.eq', ←quotient.eq', quotient_group.eq', quotient_group.eq', subgroup.mem_subgroup_of, subgroup.mem_subgroup_of, - mul_equiv.to_equiv_eq_coe, mul_equiv.coe_to_equiv, ←mul_equiv.map_inv, ←mul_equiv.map_mul, - subgroup.equiv_smul_apply_coe, subgroup.smul_mem_pointwise_smul_iff] }) + ←mul_equiv.map_inv, ←mul_equiv.map_mul, subgroup.equiv_smul_apply_coe], + exact subgroup.smul_mem_pointwise_smul_iff.symm }) lemma commensurable_conj {H K : subgroup G} (g : conj_act G) : commensurable H K ↔ commensurable (g • H) (g • K) := diff --git a/src/group_theory/subgroup/basic.lean b/src/group_theory/subgroup/basic.lean index 2cf2efd18c74e..a5c2c2982e21d 100644 --- a/src/group_theory/subgroup/basic.lean +++ b/src/group_theory/subgroup/basic.lean @@ -87,9 +87,10 @@ membership of a subgroup's underlying set. subgroup, subgroups -/ +open function open_locale big_operators -variables {G : Type*} [group G] +variables {G G' : Type*} [group G] [group G'] variables {A : Type*} [add_group A] section subgroup_class @@ -1781,16 +1782,16 @@ instance is_commutative.comm_group [h : H.is_commutative] : comm_group H := instance center.is_commutative : (center G).is_commutative := ⟨⟨λ a b, subtype.ext (b.2 a)⟩⟩ -@[to_additive] instance map_is_commutative {G' : Type*} [group G'] (f : G →* G') - [H.is_commutative] : (H.map f).is_commutative := +@[to_additive] instance map_is_commutative (f : G →* G') [H.is_commutative] : + (H.map f).is_commutative := ⟨⟨begin rintros ⟨-, a, ha, rfl⟩ ⟨-, b, hb, rfl⟩, rw [subtype.ext_iff, coe_mul, coe_mul, subtype.coe_mk, subtype.coe_mk, ←map_mul, ←map_mul], exact congr_arg f (subtype.ext_iff.mp (mul_comm ⟨a, ha⟩ ⟨b, hb⟩)), end⟩⟩ -@[to_additive] lemma comap_injective_is_commutative {G' : Type*} [group G'] {f : G' →* G} - (hf : function.injective f) [H.is_commutative] : (H.comap f).is_commutative := +@[to_additive] lemma comap_injective_is_commutative {f : G' →* G} (hf : injective f) + [H.is_commutative] : (H.comap f).is_commutative := ⟨⟨λ a b, subtype.ext begin have := mul_comm (⟨f a, a.2⟩ : H) (⟨f b, b.2⟩ : H), rwa [subtype.ext_iff, coe_mul, coe_mul, coe_mk, coe_mk, ←map_mul, ←map_mul, hf.eq_iff] at this, @@ -2405,8 +2406,10 @@ comap_sup_eq_of_le_range L.subtype (hH.trans L.subtype_range.ge) (hK.trans L.sub codisjoint (H.subgroup_of (H ⊔ K)) (K.subgroup_of (H ⊔ K)) := by { rw [codisjoint_iff, sup_subgroup_of_eq, subgroup_of_self], exacts [le_sup_left, le_sup_right] } -/-- A subgroup is isomorphic to its image under an injective function -/ -@[to_additive "An additive subgroup is isomorphic to its image under an injective function"] +/-- A subgroup is isomorphic to its image under an injective function. If you have an isomorphism, +use `mul_equiv.subgroup_map` for better definitional equalities. -/ +@[to_additive "An additive subgroup is isomorphic to its image under an injective function. If you +have an isomorphism, use `add_equiv.add_subgroup_map` for better definitional equalities."] noncomputable def equiv_map_of_injective (H : subgroup G) (f : G →* N) (hf : function.injective f) : H ≃* H.map f := { map_mul' := λ _ _, subtype.ext (f.map_mul _ _), ..equiv.set.image f H hf } @@ -2764,8 +2767,6 @@ end subgroup namespace monoid_hom -variables {G' : Type*} [group G'] - /-- The `monoid_hom` from the preimage of a subgroup to itself. -/ @[to_additive "the `add_monoid_hom` from the preimage of an additive subgroup to itself.", simps] def subgroup_comap (f : G →* G') (H' : subgroup G') : H'.comap f →* H' := @@ -2784,7 +2785,6 @@ f.submonoid_map_surjective H.to_submonoid end monoid_hom namespace mul_equiv - variables {H K : subgroup G} /-- Makes the identity isomorphism from a proof two subgroups of a multiplicative @@ -2794,20 +2794,30 @@ two subgroups of an additive group are equal."] def subgroup_congr (h : H = K) : H ≃* K := { map_mul' := λ _ _, rfl, ..equiv.set_congr $ congr_arg _ h } -/-- A `mul_equiv` `φ` between two groups `G` and `G'` induces a `mul_equiv` between -a subgroup `H ≤ G` and the subgroup `φ(H) ≤ G'`. -/ -@[to_additive "An `add_equiv` `φ` between two additive groups `G` and `G'` induces an `add_equiv` -between a subgroup `H ≤ G` and the subgroup `φ(H) ≤ G'`. "] -def subgroup_map {G'} [group G'] (e : G ≃* G') (H : subgroup G) : - H ≃* H.map e.to_monoid_hom := -e.submonoid_map H.to_submonoid +/-- A subgroup is isomorphic to its image under an isomorphism. If you only have an injective map, +use `subgroup.equiv_map_of_injective`. -/ +@[to_additive "An additive subgroup is isomorphic to its image under an an isomorphism. If you only +have an injective map, use `add_subgroup.equiv_map_of_injective`."] +def subgroup_map (e : G ≃* G') (H : subgroup G) : H ≃* H.map (e : G →* G') := +mul_equiv.submonoid_map (e : G ≃* G') H.to_submonoid -end mul_equiv +@[simp, to_additive] +lemma coe_subgroup_map_apply (e : G ≃* G') (H : subgroup G) (g : H) : + ((subgroup_map e H g : H.map (e : G →* G')) : G') = e g := rfl + +@[simp, to_additive] +lemma subgroup_map_symm_apply (e : G ≃* G') (H : subgroup G) (g : H.map (e : G →* G')) : + (e.subgroup_map H).symm g = ⟨e.symm g, set_like.mem_coe.1 $ set.mem_image_equiv.1 g.2⟩ := rfl --- TODO : ↥(⊤ : subgroup H) ≃* H ? +end mul_equiv namespace subgroup +@[simp, to_additive] +lemma equiv_map_of_injective_coe_mul_equiv (H : subgroup G) (e : G ≃* G') : + H.equiv_map_of_injective (e : G →* G') (equiv_like.injective e) = e.subgroup_map H := +by { ext, refl } + variables {C : Type*} [comm_group C] {s t : subgroup C} {x : C} @[to_additive] diff --git a/src/group_theory/submonoid/operations.lean b/src/group_theory/submonoid/operations.lean index 6dcaf235d28e3..6226d6e25dc76 100644 --- a/src/group_theory/submonoid/operations.lean +++ b/src/group_theory/submonoid/operations.lean @@ -605,8 +605,10 @@ def top_equiv : (⊤ : submonoid M) ≃* M := (top_equiv : _ ≃* M).to_monoid_hom = (⊤ : submonoid M).subtype := rfl -/-- A submonoid is isomorphic to its image under an injective function -/ -@[to_additive "An additive submonoid is isomorphic to its image under an injective function"] +/-- A subgroup is isomorphic to its image under an injective function. If you have an isomorphism, +use `mul_equiv.submonoid_map` for better definitional equalities. -/ +@[to_additive "An additive subgroup is isomorphic to its image under an injective function. If you +have an isomorphism, use `add_equiv.add_submonoid_map` for better definitional equalities."] noncomputable def equiv_map_of_injective (f : M →* N) (hf : function.injective f) : S ≃* S.map f := { map_mul' := λ _ _, subtype.ext (f.map_mul _ _), ..equiv.set.image f S hf } @@ -1055,15 +1057,25 @@ a submonoid `S ≤ M` and the submonoid `φ(S) ≤ N`. See `monoid_hom.submonoid_map` for a variant for `monoid_hom`s. -/ @[to_additive "An `add_equiv` `φ` between two additive monoids `M` and `N` induces an `add_equiv` between a submonoid `S ≤ M` and the submonoid `φ(S) ≤ N`. See `add_monoid_hom.add_submonoid_map` -for a variant for `add_monoid_hom`s.", simps] +for a variant for `add_monoid_hom`s."] def submonoid_map (e : M ≃* N) (S : submonoid M) : S ≃* S.map e.to_monoid_hom := -{ to_fun := λ x, ⟨e x, _⟩, - inv_fun := λ x, ⟨e.symm x, _⟩, -- we restate this for `simps` to avoid `⇑e.symm.to_equiv x` - ..e.to_monoid_hom.submonoid_map S, - ..e.to_equiv.image S } +{ map_mul' := λ _ _, subtype.ext (map_mul e _ _), ..(e : M ≃ N).image S } + +@[simp, to_additive] +lemma coe_submonoid_map_apply (e : M ≃* N) (S : submonoid M) (g : S) : + ((submonoid_map e S g : S.map (e : M →* N)) : N) = e g := rfl + +@[simp, to_additive add_equiv.add_submonoid_map_symm_apply] +lemma submonoid_map_symm_apply (e : M ≃* N) (S : submonoid M) (g : S.map (e : M →* N)) : + (e.submonoid_map S).symm g = ⟨e.symm g, set_like.mem_coe.1 $ set.mem_image_equiv.1 g.2⟩ := rfl end mul_equiv +@[simp, to_additive] +lemma submonoid.equiv_map_of_injective_coe_mul_equiv (e : M ≃* N) : + S.equiv_map_of_injective (e : M →* N) (equiv_like.injective e) = e.submonoid_map S := +by { ext, refl } + section actions /-! ### Actions by `submonoid`s From 12c24b7967888500356644d6250b8607c93d9c9a Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 13 Jan 2023 15:23:31 +0000 Subject: [PATCH 0232/1291] =?UTF-8?q?feat(algebra/smul=5Fwith=5Fzero):=20`?= =?UTF-8?q?a=20=E2=80=A2=20b=20=E2=89=A0=200=20=E2=86=92=20a=20=E2=89=A0?= =?UTF-8?q?=200`=20(#18086)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This matches existing `mul` lemmas --- src/algebra/smul_with_zero.lean | 7 +++++++ 1 file changed, 7 insertions(+) diff --git a/src/algebra/smul_with_zero.lean b/src/algebra/smul_with_zero.lean index fd32bf48762f0..e37fcce9ff05c 100644 --- a/src/algebra/smul_with_zero.lean +++ b/src/algebra/smul_with_zero.lean @@ -63,6 +63,13 @@ variables (R) {M} [has_zero R] [has_zero M] [smul_with_zero R M] @[simp] lemma zero_smul (m : M) : (0 : R) • m = 0 := smul_with_zero.zero_smul m +variables {R} {a : R} {b : M} + +lemma smul_eq_zero_of_left (h : a = 0) (b : M) : a • b = 0 := h.symm ▸ zero_smul _ b +lemma smul_eq_zero_of_right (a : R) (h : b = 0) : a • b = 0 := h.symm ▸ smul_zero a +lemma left_ne_zero_of_smul : a • b ≠ 0 → a ≠ 0 := mt $ λ h, smul_eq_zero_of_left h b +lemma right_ne_zero_of_smul : a • b ≠ 0 → b ≠ 0 := mt $ smul_eq_zero_of_right a + variables {R M} [has_zero R'] [has_zero M'] [has_smul R M'] /-- Pullback a `smul_with_zero` structure along an injective zero-preserving homomorphism. From 1990ff7e1ce55c2d37ad7e357976f1afb0da6618 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 13 Jan 2023 18:13:58 +0000 Subject: [PATCH 0233/1291] feat(data/finset/sups): Set family operations (#17947) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Define three binary operations on `set α` and `finset α` for use in the four functions theorem and the van den Berg-Kesten-Reimer and Ahlswede-Daykin inequalities. --- docs/references.bib | 9 ++ src/data/finset/sups.lean | 292 ++++++++++++++++++++++++++++++++++++++ 2 files changed, 301 insertions(+) create mode 100644 src/data/finset/sups.lean diff --git a/docs/references.bib b/docs/references.bib index 47a3cec0899f5..32680882acc87 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -224,6 +224,15 @@ @Article{ birkhoff1942 url = {https://doi.org/10.2307/1968871} } +@Book{ bollobas1986, + author = {Bollob\'{a}s, B\'{e}la}, + title = {Combinatorics: Set Systems, Hypergraphs, Families of Vectors, and Combinatorial + Probability}, + year = {1986}, + isbn = {0521330599}, + publisher = {Cambridge University Press} +} + @Book{ borceux-vol1, title = {Handbook of Categorical Algebra: Volume 1, Basic Category Theory}, diff --git a/src/data/finset/sups.lean b/src/data/finset/sups.lean new file mode 100644 index 0000000000000..97f93b29855e2 --- /dev/null +++ b/src/data/finset/sups.lean @@ -0,0 +1,292 @@ +/- +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.n_ary + +/-! +# Set family operations + +This file defines a few binary operations on `finset α` for use in set family combinatorics. + +## Main declarations + +* `finset.sups s t`: Finset of elements of the form `a ⊔ b` where `a ∈ s`, `b ∈ t`. +* `finset.infs s t`: Finset of elements of the form `a ⊓ b` where `a ∈ s`, `b ∈ t`. +* `finset.disj_sups s t`: Finset of elements of the form `a ⊔ b` where `a ∈ s`, `b ∈ t` and `a` + and `b` are disjoint. + +## Notation + +We define the following notation in locale `finset_family`: +* `s ⊻ t` for `finset.sups s t` +* `s ⊼ t` for `finset.infs s t` +* `s ○ t` for `finset.disj_sups s t` + +## References + +[B. Bollobás, *Combinatorics*][bollobas1986] +-/ + +open function + +variables {α : Type*} [decidable_eq α] + +namespace finset +section sups +variables [semilattice_sup α] (s s₁ s₂ t t₁ t₂ u : finset α) + +/-- The finset of elements of the form `a ⊔ b` where `a ∈ s`, `b ∈ t`. -/ +def sups (s t : finset α) : finset α := image₂ (⊔) s t + +localized "infix (name := finset.sups) ` ⊻ `:74 := finset.sups" in finset_family + +variables {s t} {a b c : α} + +@[simp] lemma mem_sups : c ∈ s ⊻ t ↔ ∃ (a ∈ s) (b ∈ t), a ⊔ b = c := by simp [sups] + +variables (s t) + +@[simp, norm_cast] lemma coe_sups : (s ⊻ t : set α) = set.image2 (⊔) s t := coe_image₂ _ _ _ + +lemma card_sups_le : (s ⊻ t).card ≤ s.card * t.card := card_image₂_le _ _ _ + +lemma card_sups_iff : + (s ⊻ t).card = s.card * t.card ↔ (s ×ˢ t : set (α × α)).inj_on (λ x, x.1 ⊔ x.2) := +card_image₂_iff + +variables {s s₁ s₂ t t₁ t₂ u} + +lemma sup_mem_sups : a ∈ s → b ∈ t → a ⊔ b ∈ s ⊻ t := mem_image₂_of_mem +lemma sups_subset : s₁ ⊆ s₂ → t₁ ⊆ t₂ → s₁ ⊻ t₁ ⊆ s₂ ⊻ t₂ := image₂_subset +lemma sups_subset_left : t₁ ⊆ t₂ → s ⊻ t₁ ⊆ s ⊻ t₂ := image₂_subset_left +lemma sups_subset_right : s₁ ⊆ s₂ → s₁ ⊻ t ⊆ s₂ ⊻ t := image₂_subset_right + +lemma image_subset_sups_left : b ∈ t → (λ a, a ⊔ b) '' s ⊆ s ⊻ t := image_subset_image₂_left +lemma image_subset_sups_right : a ∈ s → (⊔) a '' t ⊆ s ⊻ t := image_subset_image₂_right + +lemma forall_sups_iff {p : α → Prop} : (∀ c ∈ s ⊻ t, p c) ↔ ∀ (a ∈ s) (b ∈ t), p (a ⊔ b) := +forall_image₂_iff + +@[simp] lemma sups_subset_iff : s ⊻ t ⊆ u ↔ ∀ (a ∈ s) (b ∈ t), a ⊔ b ∈ u := image₂_subset_iff + +@[simp] lemma sups_nonempty_iff : (s ⊻ t).nonempty ↔ s.nonempty ∧ t.nonempty := image₂_nonempty_iff + +lemma nonempty.sups : s.nonempty → t.nonempty → (s ⊻ t).nonempty := nonempty.image₂ +lemma nonempty.of_sups_left : (s ⊻ t).nonempty → s.nonempty := nonempty.of_image₂_left +lemma nonempty.of_sups_right : (s ⊻ t).nonempty → t.nonempty := nonempty.of_image₂_right + +@[simp] lemma sups_empty_left : ∅ ⊻ t = ∅ := image₂_empty_left +@[simp] lemma sups_empty_right : s ⊻ ∅ = ∅ := image₂_empty_right +@[simp] lemma sups_eq_empty_iff : s ⊻ t = ∅ ↔ s = ∅ ∨ t = ∅ := image₂_eq_empty_iff + +@[simp] lemma sups_singleton_left : {a} ⊻ t = t.image (λ b, a ⊔ b) := image₂_singleton_left +@[simp] lemma sups_singleton_right : s ⊻ {b} = s.image (λ a, a ⊔ b) := image₂_singleton_right +lemma sups_singleton_left' : {a} ⊻ t = t.image ((⊔) a) := image₂_singleton_left' + +lemma sups_singleton : ({a} ⊻ {b} : finset α) = {a ⊔ b} := image₂_singleton + +lemma sups_union_left : (s₁ ∪ s₂) ⊻ t = s₁ ⊻ t ∪ s₂ ⊻ t := image₂_union_left +lemma sups_union_right : s ⊻ (t₁ ∪ t₂) = s ⊻ t₁ ∪ s ⊻ t₂ := image₂_union_right + +lemma sups_inter_subset_left : (s₁ ∩ s₂) ⊻ t ⊆ s₁ ⊻ t ∩ s₂ ⊻ t := image₂_inter_subset_left +lemma sups_inter_subset_right : s ⊻ (t₁ ∩ t₂) ⊆ s ⊻ t₁ ∩ s ⊻ t₂ := image₂_inter_subset_right + +lemma subset_sups {s t : set α} : + ↑u ⊆ set.image2 (⊔) s t → ∃ s' t' : finset α, ↑s' ⊆ s ∧ ↑t' ⊆ t ∧ u ⊆ s' ⊻ t' := +subset_image₂ + +variables (s t u) + +lemma bUnion_image_sup_left : s.bUnion (λ a, t.image $ (⊔) a) = s ⊻ t := bUnion_image_left +lemma bUnion_image_sup_right : t.bUnion (λ b, s.image $ λ a, a ⊔ b) = s ⊻ t := bUnion_image_right + +@[simp] lemma image_sup_product (s t : finset α) : (s ×ˢ t).image (uncurry (⊔)) = s ⊻ t := +image_uncurry_product _ _ _ + +lemma sups_assoc : (s ⊻ t) ⊻ u = s ⊻ (t ⊻ u) := image₂_assoc $ λ _ _ _, sup_assoc +lemma sups_comm : s ⊻ t = t ⊻ s := image₂_comm $ λ _ _, sup_comm +lemma sups_left_comm : s ⊻ (t ⊻ u) = t ⊻ (s ⊻ u) := image₂_left_comm sup_left_comm +lemma sups_right_comm : (s ⊻ t) ⊻ u = (s ⊻ u) ⊻ t := image₂_right_comm sup_right_comm + +end sups + +section infs +variables [semilattice_inf α] (s s₁ s₂ t t₁ t₂ u : finset α) + +/-- The finset of elements of the form `a ⊓ b` where `a ∈ s`, `b ∈ t`. -/ +def infs (s t : finset α) : finset α := image₂ (⊓) s t + +localized "infix (name := finset.infs) ` ⊼ `:74 := finset.infs" in finset_family + +variables {s t} {a b c : α} + +@[simp] lemma mem_infs : c ∈ s ⊼ t ↔ ∃ (a ∈ s) (b ∈ t), a ⊓ b = c := by simp [infs] + +variables (s t) + +@[simp, norm_cast] lemma coe_infs : (s ⊼ t : set α) = set.image2 (⊓) s t := coe_image₂ _ _ _ + +lemma card_infs_le : (s ⊼ t).card ≤ s.card * t.card := card_image₂_le _ _ _ + +lemma card_infs_iff : + (s ⊼ t).card = s.card * t.card ↔ (s ×ˢ t : set (α × α)).inj_on (λ x, x.1 ⊓ x.2) := +card_image₂_iff + +variables {s s₁ s₂ t t₁ t₂ u} + +lemma inf_mem_infs : a ∈ s → b ∈ t → a ⊓ b ∈ s ⊼ t := mem_image₂_of_mem +lemma infs_subset : s₁ ⊆ s₂ → t₁ ⊆ t₂ → s₁ ⊼ t₁ ⊆ s₂ ⊼ t₂ := image₂_subset +lemma infs_subset_left : t₁ ⊆ t₂ → s ⊼ t₁ ⊆ s ⊼ t₂ := image₂_subset_left +lemma infs_subset_right : s₁ ⊆ s₂ → s₁ ⊼ t ⊆ s₂ ⊼ t := image₂_subset_right + +lemma image_subset_infs_left : b ∈ t → (λ a, a ⊓ b) '' s ⊆ s ⊼ t := image_subset_image₂_left +lemma image_subset_infs_right : a ∈ s → (⊓) a '' t ⊆ s ⊼ t := image_subset_image₂_right + +lemma forall_infs_iff {p : α → Prop} : (∀ c ∈ s ⊼ t, p c) ↔ ∀ (a ∈ s) (b ∈ t), p (a ⊓ b) := +forall_image₂_iff + +@[simp] lemma infs_subset_iff : s ⊼ t ⊆ u ↔ ∀ (a ∈ s) (b ∈ t), a ⊓ b ∈ u := image₂_subset_iff + +@[simp] lemma infs_nonempty_iff : (s ⊼ t).nonempty ↔ s.nonempty ∧ t.nonempty := image₂_nonempty_iff + +lemma nonempty.infs : s.nonempty → t.nonempty → (s ⊼ t).nonempty := nonempty.image₂ +lemma nonempty.of_infs_left : (s ⊼ t).nonempty → s.nonempty := nonempty.of_image₂_left +lemma nonempty.of_infs_right : (s ⊼ t).nonempty → t.nonempty := nonempty.of_image₂_right + +@[simp] lemma infs_empty_left : ∅ ⊼ t = ∅ := image₂_empty_left +@[simp] lemma infs_empty_right : s ⊼ ∅ = ∅ := image₂_empty_right +@[simp] lemma infs_eq_empty_iff : s ⊼ t = ∅ ↔ s = ∅ ∨ t = ∅ := image₂_eq_empty_iff + +@[simp] lemma infs_singleton_left : {a} ⊼ t = t.image (λ b, a ⊓ b) := image₂_singleton_left +@[simp] lemma infs_singleton_right : s ⊼ {b} = s.image (λ a, a ⊓ b) := image₂_singleton_right +lemma infs_singleton_left' : {a} ⊼ t = t.image ((⊓) a) := image₂_singleton_left' + +lemma infs_singleton : ({a} ⊼ {b} : finset α) = {a ⊓ b} := image₂_singleton + +lemma infs_union_left : (s₁ ∪ s₂) ⊼ t = s₁ ⊼ t ∪ s₂ ⊼ t := image₂_union_left +lemma infs_union_right : s ⊼ (t₁ ∪ t₂) = s ⊼ t₁ ∪ s ⊼ t₂ := image₂_union_right + +lemma infs_inter_subset_left : (s₁ ∩ s₂) ⊼ t ⊆ s₁ ⊼ t ∩ s₂ ⊼ t := image₂_inter_subset_left +lemma infs_inter_subset_right : s ⊼ (t₁ ∩ t₂) ⊆ s ⊼ t₁ ∩ s ⊼ t₂ := image₂_inter_subset_right + +lemma subset_infs {s t : set α} : + ↑u ⊆ set.image2 (⊓) s t → ∃ s' t' : finset α, ↑s' ⊆ s ∧ ↑t' ⊆ t ∧ u ⊆ s' ⊼ t' := +subset_image₂ + +variables (s t u) + +lemma bUnion_image_inf_left : s.bUnion (λ a, t.image $ (⊓) a) = s ⊼ t := bUnion_image_left +lemma bUnion_image_inf_right : t.bUnion (λ b, s.image $ λ a, a ⊓ b) = s ⊼ t := bUnion_image_right + +@[simp] lemma image_inf_product (s t : finset α) : (s ×ˢ t).image (uncurry (⊓)) = s ⊼ t := +image_uncurry_product _ _ _ + +lemma infs_assoc : (s ⊼ t) ⊼ u = s ⊼ (t ⊼ u) := image₂_assoc $ λ _ _ _, inf_assoc +lemma infs_comm : s ⊼ t = t ⊼ s := image₂_comm $ λ _ _, inf_comm +lemma infs_left_comm : s ⊼ (t ⊼ u) = t ⊼ (s ⊼ u) := image₂_left_comm inf_left_comm +lemma infs_right_comm : (s ⊼ t) ⊼ u = (s ⊼ u) ⊼ t := image₂_right_comm inf_right_comm + +end infs + +open_locale finset_family + +section disj_sups +variables [semilattice_sup α] [order_bot α] [@decidable_rel α disjoint] + (s s₁ s₂ t t₁ t₂ u : finset α) + +/-- The finset of elements of the form `a ⊔ b` where `a ∈ s`, `b ∈ t` and `a` and `b` are disjoint. +-/ +def disj_sups : finset α := +((s ×ˢ t).filter $ λ ab : α × α, disjoint ab.1 ab.2).image $ λ ab, ab.1 ⊔ ab.2 + +localized "infix (name := finset.disj_sups) ` ○ `:74 := finset.disj_sups" in finset_family + +variables {s t u} {a b c : α} + +@[simp] lemma mem_disj_sups : c ∈ s ○ t ↔ ∃ (a ∈ s) (b ∈ t), disjoint a b ∧ a ⊔ b = c := +by simp [disj_sups, and_assoc] + +lemma disj_sups_subset_sups : s ○ t ⊆ s ⊻ t := +begin + simp_rw [subset_iff, mem_sups, mem_disj_sups], + exact λ c ⟨a, b, ha, hb, h, hc⟩, ⟨a, b, ha, hb, hc⟩, +end + +variables (s t) + +lemma card_disj_sups_le : (s ○ t).card ≤ s.card * t.card := +(card_le_of_subset disj_sups_subset_sups).trans $ card_sups_le _ _ + +variables {s s₁ s₂ t t₁ t₂ u} + +lemma disj_sups_subset (hs : s₁ ⊆ s₂) (ht : t₁ ⊆ t₂) : s₁ ○ t₁ ⊆ s₂ ○ t₂ := +image_subset_image $ filter_subset_filter _ $ product_subset_product hs ht + +lemma disj_sups_subset_left (ht : t₁ ⊆ t₂) : s ○ t₁ ⊆ s ○ t₂ := disj_sups_subset subset.rfl ht +lemma disj_sups_subset_right (hs : s₁ ⊆ s₂) : s₁ ○ t ⊆ s₂ ○ t := disj_sups_subset hs subset.rfl + +lemma forall_disj_sups_iff {p : α → Prop} : + (∀ c ∈ s ○ t, p c) ↔ ∀ (a ∈ s) (b ∈ t), disjoint a b → p (a ⊔ b) := +begin + simp_rw mem_disj_sups, + refine ⟨λ h a ha b hb hab, h _ ⟨_, ha, _, hb, hab, rfl⟩, _⟩, + rintro h _ ⟨a, ha, b, hb, hab, rfl⟩, + exact h _ ha _ hb hab, +end + +@[simp] lemma disj_sups_subset_iff : s ○ t ⊆ u ↔ ∀ (a ∈ s) (b ∈ t), disjoint a b → a ⊔ b ∈ u := +forall_disj_sups_iff + +lemma nonempty.of_disj_sups_left : (s ○ t).nonempty → s.nonempty := +by { simp_rw [finset.nonempty, mem_disj_sups], exact λ ⟨_, a, ha, _⟩, ⟨a, ha⟩ } + +lemma nonempty.of_disj_sups_right : (s ○ t).nonempty → t.nonempty := +by { simp_rw [finset.nonempty, mem_disj_sups], exact λ ⟨_, _, _, b, hb, _⟩, ⟨b, hb⟩ } + +@[simp] lemma disj_sups_empty_left : ∅ ○ t = ∅ := by simp [disj_sups] +@[simp] lemma disj_sups_empty_right : s ○ ∅ = ∅ := by simp [disj_sups] + +lemma disj_sups_singleton : ({a} ○ {b} : finset α) = if disjoint a b then {a ⊔ b} else ∅ := +by split_ifs; simp [disj_sups, filter_singleton, h] + +lemma disj_sups_union_left : (s₁ ∪ s₂) ○ t = s₁ ○ t ∪ s₂ ○ t := +by simp [disj_sups, filter_union, image_union] +lemma disj_sups_union_right : s ○ (t₁ ∪ t₂) = s ○ t₁ ∪ s ○ t₂ := +by simp [disj_sups, filter_union, image_union] + +lemma disj_sups_inter_subset_left : (s₁ ∩ s₂) ○ t ⊆ s₁ ○ t ∩ s₂ ○ t := +by simpa only [disj_sups, inter_product, filter_inter_distrib] using image_inter_subset _ _ _ +lemma disj_sups_inter_subset_right : s ○ (t₁ ∩ t₂) ⊆ s ○ t₁ ∩ s ○ t₂ := +by simpa only [disj_sups, product_inter, filter_inter_distrib] using image_inter_subset _ _ _ + +variables (s t) + +lemma disj_sups_comm : s ○ t = t ○ s := +by { ext, rw [mem_disj_sups, exists₂_comm], simp [sup_comm, disjoint.comm] } + +end disj_sups + +open_locale finset_family + +section distrib_lattice +variables [distrib_lattice α] [order_bot α] [@decidable_rel α disjoint] (s t u : finset α) + +lemma disj_sups_assoc : ∀ s t u : finset α, (s ○ t) ○ u = s ○ (t ○ u) := +begin + refine associative_of_commutative_of_le disj_sups_comm _, + simp only [le_eq_subset, disj_sups_subset_iff, mem_disj_sups], + rintro s t u _ ⟨a, ha, b, hb, hab, rfl⟩ c hc habc, + rw disjoint_sup_left at habc, + exact ⟨a, ha, _, ⟨b, hb, c, hc, habc.2, rfl⟩, hab.sup_right habc.1, sup_assoc.symm⟩, +end + +lemma disj_sups_left_comm : s ○ (t ○ u) = t ○ (s ○ u) := +by simp_rw [←disj_sups_assoc, disj_sups_comm s] + +lemma disj_sups_right_comm : (s ○ t) ○ u = (s ○ u) ○ t := +by simp_rw [disj_sups_assoc, disj_sups_comm] + +end distrib_lattice +end finset From 8cce17e5783303db93df6810de9d85dd0f9e402a Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Fri, 13 Jan 2023 18:13:59 +0000 Subject: [PATCH 0234/1291] feat(number_theory/zeta_values): evaluate critical values of Riemann zeta (#17994) Evaluate critical values of Riemann zeta, via pointwise convergence of Fourier series of the Bernoulli polynomials. Co-authored-by: sgouezel --- docs/100.yaml | 6 +- src/analysis/fourier/add_circle.lean | 26 +++ src/analysis/p_series.lean | 11 ++ src/number_theory/zeta_values.lean | 246 ++++++++++++++++++++++++- src/topology/algebra/infinite_sum.lean | 75 ++++++-- 5 files changed, 339 insertions(+), 25 deletions(-) diff --git a/docs/100.yaml b/docs/100.yaml index a14251360748f..53980901664be 100644 --- a/docs/100.yaml +++ b/docs/100.yaml @@ -45,10 +45,8 @@ title : Polyhedron Formula 14: title : Euler’s Summation of 1 + (1/2)^2 + (1/3)^2 + …. - author : Marc Masdeu - links : - result : https://github.com/mmasdeu/euler/blob/main/src/euler.lean#L712 - website: https://github.com/mmasdeu/euler + decl : has_sum_zeta_two + author : Marc Masdeu, David Loeffler 15: title : Fundamental Theorem of Integral Calculus decls : diff --git a/src/analysis/fourier/add_circle.lean b/src/analysis/fourier/add_circle.lean index fa8ed0c38f94b..4ad7a556d91e6 100644 --- a/src/analysis/fourier/add_circle.lean +++ b/src/analysis/fourier/add_circle.lean @@ -358,6 +358,32 @@ lemma fourier_coeff_on.const_mul {a b : ℝ} (f : ℝ → ℂ) (c : ℂ) (n : fourier_coeff_on hab (λ x, c * f x) n = c * fourier_coeff_on hab f n := fourier_coeff_on.const_smul _ _ _ _ +include hT + +lemma fourier_coeff_lift_Ioc_eq {a : ℝ} (f : ℝ → ℂ) (n : ℤ) : + fourier_coeff (add_circle.lift_Ioc T a f) n = + fourier_coeff_on (lt_add_of_pos_right a hT.out) f n := +begin + rw [fourier_coeff_on_eq_integral, fourier_coeff_eq_interval_integral, add_sub_cancel' a T], + congr' 1, + refine interval_integral.integral_congr_ae (ae_of_all _ (λ x hx, _)), + rw lift_Ioc_coe_apply, + rwa uIoc_of_le (lt_add_of_pos_right a hT.out).le at hx, +end + +lemma fourier_coeff_lift_Ico_eq {a : ℝ} (f : ℝ → ℂ) (n : ℤ) : + fourier_coeff (add_circle.lift_Ico T a f) n = + fourier_coeff_on (lt_add_of_pos_right a hT.out) f n := +begin + rw [fourier_coeff_on_eq_integral, fourier_coeff_eq_interval_integral _ _ a, add_sub_cancel' a T], + congr' 1, + simp_rw [interval_integral.integral_of_le (lt_add_of_pos_right a hT.out).le, + integral_Ioc_eq_integral_Ioo], + refine set_integral_congr measurable_set_Ioo (λ x hx, _), + dsimp only, + rw lift_Ico_coe_apply (Ioo_subset_Ico_self hx), +end + end fourier_coeff section fourier_L2 diff --git a/src/analysis/p_series.lean b/src/analysis/p_series.lean index 3698e6084e96d..d5feac20069b5 100644 --- a/src/analysis/p_series.lean +++ b/src/analysis/p_series.lean @@ -203,6 +203,17 @@ if and only if `1 < p`. -/ lemma real.summable_one_div_nat_pow {p : ℕ} : summable (λ n, 1 / n ^ p : ℕ → ℝ) ↔ 1 < p := by simp +/-- Summability of the `p`-series over `ℤ`. -/ +lemma real.summable_one_div_int_pow {p : ℕ} : summable (λ n:ℤ, 1 / (n : ℝ) ^ p) ↔ 1 < p := +begin + refine ⟨λ h, real.summable_one_div_nat_pow.mp (h.comp_injective nat.cast_injective), + λ h, summable_int_of_summable_nat (real.summable_one_div_nat_pow.mpr h) + (((real.summable_one_div_nat_pow.mpr h).mul_left $ 1 / (-1) ^ p).congr $ λ n, _)⟩, + conv_rhs { rw [int.cast_neg, neg_eq_neg_one_mul, mul_pow, ←div_div] }, + conv_lhs { rw [mul_div, mul_one], }, + refl, +end + /-- Harmonic series is not unconditionally summable. -/ lemma real.not_summable_nat_cast_inv : ¬summable (λ n, n⁻¹ : ℕ → ℝ) := have ¬summable (λ n, (n^1)⁻¹ : ℕ → ℝ), from mt real.summable_nat_pow_inv.1 (lt_irrefl 1), diff --git a/src/number_theory/zeta_values.lean b/src/number_theory/zeta_values.lean index 6484074af7da9..9f3340f03eb37 100644 --- a/src/number_theory/zeta_values.lean +++ b/src/number_theory/zeta_values.lean @@ -8,19 +8,34 @@ import number_theory.bernoulli_polynomials import analysis.special_functions.integrals import measure_theory.integral.interval_integral import analysis.fourier.add_circle +import analysis.p_series + /-! # Critical values of the Riemann zeta function -In this file we evaluate the Fourier coefficients of Bernoulli polynomials on the interval `[0, 1]`. -In a future iteration this will be used to evaluate critical values of the Riemann zeta function -(and other Dirichlet L-functions). +In this file we prove formulae for the critical values of `ζ(s)`, and more generally of Hurwitz +zeta functions, in terms of Bernoulli polynomials. + +## Main results: + +* `has_sum_zeta_nat`: the final formula for zeta values, + $$\zeta(2k) = \frac{(-1)^{(k + 1)} 2 ^ {2k - 1} \pi^{2k} B_{2 k}}{(2 k)!}.$$ +* `has_sum_zeta_two` and `has_sum_zeta_four`: special cases given explicitly. +* `has_sum_one_div_nat_pow_mul_cos`: a formula for the sum `∑ (n : ℕ), cos (2 π i n x) / n ^ k` as + an explicit multiple of `Bₖ(x)`, for any `x ∈ [0, 1]` and `k ≥ 2` even. +* `has_sum_one_div_nat_pow_mul_sin`: a formula for the sum `∑ (n : ℕ), sin (2 π i n x) / n ^ k` as + an explicit multiple of `Bₖ(x)`, for any `x ∈ [0, 1]` and `k ≥ 3` odd. -/ + noncomputable theory open_locale nat real interval open complex measure_theory set interval_integral +local notation `𝕌` := unit_add_circle +local attribute [instance] real.fact_zero_lt_one + section bernoulli_fun_props /-! Simple properties of the Bernoulli polynomial, as a function `ℝ → ℝ`. -/ @@ -78,8 +93,6 @@ end bernoulli_fun_props section bernoulli_fourier_coeffs /-! Compute the Fourier coefficients of the Bernoulli functions via integration by parts. -/ -local attribute [instance] real.fact_zero_lt_one - /-- The `n`-th Fourier coefficient of the `k`-th Bernoulli function on the interval `[0, 1]`. -/ def bernoulli_fourier_coeff (k : ℕ) (n : ℤ) : ℂ := fourier_coeff_on zero_lt_one (λ x, bernoulli_fun k x) n @@ -134,3 +147,226 @@ begin end end bernoulli_fourier_coeffs + +section bernoulli_periodized +/-! In this section we use the above evaluations of the Fourier coefficients of Bernoulli +polynomials, together with the theorem `has_pointwise_sum_fourier_series_of_summable` from Fourier +theory, to obtain an explicit formula for `∑ (n:ℤ), 1 / n ^ k * fourier n x`. -/ + +/-- The Bernoulli polynomial, extended from `[0, 1)` to the unit circle. -/ +def periodized_bernoulli (k : ℕ) : 𝕌 → ℝ := add_circle.lift_Ico 1 0 (bernoulli_fun k) + +lemma periodized_bernoulli.continuous {k : ℕ} (hk : k ≠ 1) : continuous (periodized_bernoulli k) := +add_circle.lift_Ico_zero_continuous + (by exact_mod_cast (bernoulli_fun_endpoints_eq_of_ne_one hk).symm) + (polynomial.continuous _).continuous_on + +lemma fourier_coeff_bernoulli_eq {k : ℕ} (hk : k ≠ 0) (n : ℤ) : + fourier_coeff (coe ∘ periodized_bernoulli k : 𝕌 → ℂ) n = -k! / (2 * π * I * n) ^ k := +begin + have : (coe ∘ periodized_bernoulli k : 𝕌 → ℂ) = add_circle.lift_Ico 1 0 (coe ∘ bernoulli_fun k), + { ext1 x, refl }, + rw [this, fourier_coeff_lift_Ico_eq], + simpa only [zero_add] using bernoulli_fourier_coeff_eq hk n, +end + +lemma summable_bernoulli_fourier {k : ℕ} (hk : 2 ≤ k) : + summable (λ n, -k! / (2 * π * I * n) ^ k : ℤ → ℂ) := +begin + have : ∀ (n : ℤ), -(k! : ℂ) / (2 * π * I * n) ^ k + = (-k! / (2 * π * I) ^ k) * (1 / n ^ k), + { intro n, rw [mul_one_div, div_div, ←mul_pow], }, + simp_rw this, + apply summable.mul_left, + rw ←summable_norm_iff, + have : (λ (x : ℤ), ‖1 / (x:ℂ) ^ k‖) = (λ (x : ℤ), |1 / (x:ℝ) ^ k|), + { ext1 x, + rw [norm_eq_abs, ←complex.abs_of_real], + congr' 1, + norm_cast }, + simp_rw this, + rw [summable_abs_iff], + exact real.summable_one_div_int_pow.mpr (one_lt_two.trans_le hk), +end + +lemma has_sum_one_div_pow_mul_fourier_mul_bernoulli_fun {k : ℕ} (hk : 2 ≤ k) + {x : ℝ} (hx : x ∈ Icc (0:ℝ) 1) : + has_sum (λ n:ℤ, 1 / (n:ℂ) ^ k * fourier n (x : 𝕌)) (-(2 * π * I) ^ k / k! * bernoulli_fun k x) := +begin + -- first show it suffices to prove result for `Ico 0 1` + suffices : ∀ {y : ℝ}, y ∈ Ico (0:ℝ) 1 → has_sum _ _, + { rw [←Ico_insert_right (zero_le_one' ℝ), mem_insert_iff, or.comm] at hx, + rcases hx with hx | rfl, + { exact this hx }, + { convert this (left_mem_Ico.mpr zero_lt_one) using 1, + { rw [add_circle.coe_period, quotient_add_group.coe_zero], }, + { rw bernoulli_fun_endpoints_eq_of_ne_one (by linarith : k ≠ 1) } } }, + intros y hy, + let B : C(𝕌, ℂ) := continuous_map.mk (coe ∘ periodized_bernoulli k) + (continuous_of_real.comp (periodized_bernoulli.continuous (by linarith))), + have step1 : ∀ (n:ℤ), fourier_coeff B n = -k! / (2 * π * I * n) ^ k, + { rw continuous_map.coe_mk, exact fourier_coeff_bernoulli_eq (by linarith : k ≠ 0) }, + have step2 := has_pointwise_sum_fourier_series_of_summable ((summable_bernoulli_fourier hk).congr + (λ n, (step1 n).symm)) y, + simp_rw step1 at step2, + convert step2.mul_left ((-(2 * ↑π * I) ^ k) / (k! : ℂ)) using 2, + ext1 n, + rw [smul_eq_mul, ←mul_assoc, mul_div, mul_neg, div_mul_cancel, neg_neg, mul_pow _ ↑n, ←div_div, + div_self], + { rw [ne.def, pow_eq_zero_iff', not_and_distrib], + exact or.inl two_pi_I_ne_zero, }, + { exact nat.cast_ne_zero.mpr (nat.factorial_ne_zero _), }, + { rw [continuous_map.coe_mk, function.comp_app, of_real_inj, + periodized_bernoulli, add_circle.lift_Ico_coe_apply (by rwa zero_add)] }, +end + +end bernoulli_periodized + +section cleanup +/- This section is just reformulating the results in a nicer form. -/ + +lemma has_sum_one_div_nat_pow_mul_fourier {k : ℕ} (hk : 2 ≤ k) {x : ℝ} (hx : x ∈ Icc (0:ℝ) 1) : + has_sum (λ n:ℕ, 1 / (n:ℂ) ^ k * (fourier n (x : 𝕌) + (-1) ^ k * fourier (-n) (x : 𝕌))) + (-(2 * π * I) ^ k / k! * bernoulli_fun k x) := +begin + convert (has_sum_one_div_pow_mul_fourier_mul_bernoulli_fun hk hx).sum_nat_of_sum_int, + { ext1 n, + rw [int.cast_neg, mul_add, ←mul_assoc], + conv_rhs { rw [neg_eq_neg_one_mul, mul_pow, ←div_div] }, + congr' 2, + rw [div_mul_eq_mul_div₀, one_mul], + congr' 1, + rw [eq_div_iff, ←mul_pow, ←neg_eq_neg_one_mul, neg_neg, one_pow], + apply pow_ne_zero, rw neg_ne_zero, exact one_ne_zero, }, + { rw [int.cast_zero, zero_pow (by linarith : 0 < k), div_zero, zero_mul, add_zero] }, +end + +lemma has_sum_one_div_nat_pow_mul_cos {k : ℕ} (hk : k ≠ 0) {x : ℝ} (hx : x ∈ Icc (0:ℝ) 1) : + has_sum (λ n:ℕ, 1 / (n:ℝ) ^ (2 * k) * real.cos (2 * π * n * x)) + ((-1) ^ (k + 1) * (2 * π) ^ (2 * k) / 2 / (2 * k)! * + (polynomial.map (algebra_map ℚ ℝ) (polynomial.bernoulli (2 * k))).eval x) := +begin + have : has_sum (λ n:ℕ, 1 / (n:ℂ) ^ (2 * k) * (fourier n (x : 𝕌) + fourier (-n) (x : 𝕌))) + ((-1) ^ (k + 1) * (2 * π) ^ (2 * k) / (2 * k)! * bernoulli_fun (2 * k) x), + { convert (has_sum_one_div_nat_pow_mul_fourier + (by linarith [nat.one_le_iff_ne_zero.mpr hk] : 2 ≤ 2 * k) hx), + { ext1 n, + rw [pow_mul (-1 : ℂ),neg_one_sq, one_pow, one_mul], }, + { rw [pow_add, pow_one], + conv_rhs { rw [mul_pow], congr, congr, skip, rw [pow_mul, I_sq] }, + ring, } }, + convert ((has_sum_iff _ _).mp (this.div_const 2)).1, + { ext1 n, + convert (of_real_re _).symm, + rw of_real_mul,rw ←mul_div, congr, + { rw [of_real_div, of_real_one, of_real_pow], refl, }, + { rw [of_real_cos, of_real_mul, fourier_coe_apply, fourier_coe_apply, cos, of_real_one, div_one, + div_one, of_real_mul, of_real_mul, of_real_bit0, of_real_one, int.cast_neg, + int.cast_coe_nat, of_real_nat_cast], + congr' 3, + { ring }, { ring }, }, }, + { convert (of_real_re _).symm, + rw [of_real_mul, of_real_div, of_real_div, of_real_mul, of_real_pow, of_real_pow, of_real_neg, + of_real_nat_cast, of_real_mul, of_real_bit0, of_real_one], + ring }, +end + +lemma has_sum_one_div_nat_pow_mul_sin {k : ℕ} (hk : k ≠ 0) {x : ℝ} (hx : x ∈ Icc (0:ℝ) 1) : + has_sum (λ n:ℕ, 1 / (n:ℝ) ^ (2 * k + 1) * real.sin (2 * π * n * x)) + ((-1) ^ (k + 1) * (2 * π) ^ (2 * k + 1) / 2 / (2 * k + 1)! * + (polynomial.map (algebra_map ℚ ℝ) (polynomial.bernoulli (2 * k + 1))).eval x) := +begin + have : has_sum (λ n:ℕ, 1 / (n:ℂ) ^ (2 * k + 1) * (fourier n (x : 𝕌) - fourier (-n) (x : 𝕌))) + ((-1)^(k + 1) * I * (2 * π)^(2 * k + 1) / (2 * k + 1)! * bernoulli_fun (2 * k + 1) x), + { convert (has_sum_one_div_nat_pow_mul_fourier + (by linarith [nat.one_le_iff_ne_zero.mpr hk] : 2 ≤ 2 * k + 1) hx), + { ext1 n, + rw [pow_add (-1: ℂ), pow_mul (-1 : ℂ), neg_one_sq, one_pow, one_mul, pow_one, + ←neg_eq_neg_one_mul, ←sub_eq_add_neg], }, + { rw [pow_add, pow_one], + conv_rhs { rw [mul_pow], congr, congr, skip, rw [pow_add, pow_one, pow_mul, I_sq] }, + ring, }, }, + convert ((has_sum_iff _ _).mp (this.div_const (2 * I))).1, + { ext1 n, + convert (of_real_re _).symm, + rw of_real_mul,rw ←mul_div, congr, + { rw [of_real_div, of_real_one, of_real_pow], refl, }, + { rw [of_real_sin, of_real_mul, fourier_coe_apply, fourier_coe_apply, sin, of_real_one, div_one, + div_one, of_real_mul, of_real_mul, of_real_bit0, of_real_one, int.cast_neg, + int.cast_coe_nat, of_real_nat_cast, ←div_div, div_I, div_mul_eq_mul_div₀, ←neg_div, + ←neg_mul, neg_sub], + congr' 4, + { ring, }, { ring }, }, }, + { convert (of_real_re _).symm, + rw [of_real_mul, of_real_div, of_real_div, of_real_mul, of_real_pow, of_real_pow, of_real_neg, + of_real_nat_cast, of_real_mul, of_real_bit0, of_real_one, + ←div_div, div_I, div_mul_eq_mul_div₀], + have : ∀ (α β γ δ : ℂ), α * I * β / γ * δ * I = I ^ 2 * α * β / γ * δ := by { intros, ring }, + rw [this, I_sq], + ring, }, +end + +lemma has_sum_zeta_nat {k : ℕ} (hk : k ≠ 0) : has_sum (λ n:ℕ, 1 / (n:ℝ) ^ (2 * k)) + ((-1) ^ (k + 1) * 2 ^ (2 * k - 1) * π ^ (2 * k) * bernoulli (2 * k) / (2 * k)!) := +begin + convert has_sum_one_div_nat_pow_mul_cos hk (left_mem_Icc.mpr zero_le_one), + { ext1 n, rw [mul_zero, real.cos_zero, mul_one], }, + rw [polynomial.eval_zero_map, polynomial.bernoulli_eval_zero, eq_rat_cast], + have : (2:ℝ) ^ (2 * k - 1) = (2:ℝ) ^ (2 * k) / 2, + { rw eq_div_iff (two_ne_zero' ℝ), + conv_lhs { congr, skip, rw ←pow_one (2:ℝ) }, + rw [←pow_add, nat.sub_add_cancel], + linarith [nat.one_le_iff_ne_zero.mpr hk], }, + rw [this, mul_pow], + ring, +end + +end cleanup + +section examples + +lemma has_sum_zeta_two : has_sum (λ n:ℕ, 1 / (n : ℝ) ^ 2) (π ^ 2 / 6) := +begin + convert has_sum_zeta_nat one_ne_zero using 1, rw mul_one, + rw [bernoulli_eq_bernoulli'_of_ne_one (by dec_trivial : 2 ≠ 1), bernoulli'_two], + norm_num, field_simp, ring, +end + +lemma has_sum_zeta_four : has_sum (λ n:ℕ, 1 / (n : ℝ) ^ 4) (π ^ 4 / 90) := +begin + convert has_sum_zeta_nat two_ne_zero using 1, norm_num, + rw [bernoulli_eq_bernoulli'_of_ne_one, bernoulli'_four], + norm_num, field_simp, ring, dec_trivial, +end + +lemma polynomial.bernoulli_three_eval_one_quarter : + (polynomial.bernoulli 3).eval (1 / 4) = 3 / 64 := +begin + simp_rw [polynomial.bernoulli, finset.sum_range_succ, polynomial.eval_add, + polynomial.eval_monomial], + rw [finset.sum_range_zero, polynomial.eval_zero, zero_add, bernoulli_one], + rw [bernoulli_eq_bernoulli'_of_ne_one zero_ne_one, bernoulli'_zero, + bernoulli_eq_bernoulli'_of_ne_one (by dec_trivial : 2 ≠ 1), bernoulli'_two, + bernoulli_eq_bernoulli'_of_ne_one (by dec_trivial : 3 ≠ 1), bernoulli'_three], + norm_num, +end + +/-- Explicit formula for `L(χ, 3)`, where `χ` is the unique nontrivial Dirichlet character modulo 4. +-/ +lemma has_sum_L_function_mod_four_eval_three : + has_sum (λ n:ℕ, (1 / (n:ℝ) ^ 3 * real.sin (π * n / 2))) (π ^ 3 / 32) := +begin + convert has_sum_one_div_nat_pow_mul_sin one_ne_zero (_ : 1 / 4 ∈ Icc (0:ℝ) 1), + { ext1 n, + norm_num, + left, + congr' 1, + ring, }, + { have : (1 / 4 : ℝ) = (algebra_map ℚ ℝ) (1 / 4 : ℚ), by norm_num, + rw [this, mul_pow, polynomial.eval_map, polynomial.eval₂_at_apply, + (by dec_trivial : 2 * 1 + 1 = 3), polynomial.bernoulli_three_eval_one_quarter], + norm_num, field_simp, ring }, + { rw mem_Icc, split, linarith, linarith, }, +end + +end examples diff --git a/src/topology/algebra/infinite_sum.lean b/src/topology/algebra/infinite_sum.lean index 885d26f843fd5..2ba8e2083616a 100644 --- a/src/topology/algebra/infinite_sum.lean +++ b/src/topology/algebra/infinite_sum.lean @@ -836,13 +836,12 @@ end end tsum /-! -### Sums on subtypes +### Sums on nat -If `s` is a finset of `α`, we show that the summability of `f` in the whole space and on the subtype -`univ - s` are equivalent, and relate their sums. For a function defined on `ℕ`, we deduce the -formula `(∑ i in range k, f i) + (∑' i, f (i + k)) = (∑' i, f i)`, in `sum_add_tsum_nat_add`. +We show the formula `(∑ i in range k, f i) + (∑' i, f (i + k)) = (∑' i, f i)`, in +`sum_add_tsum_nat_add`, as well as several results relating sums on `ℕ` and `ℤ`. -/ -section subtype +section nat lemma has_sum_nat_add_iff {f : ℕ → α} (k : ℕ) {a : α} : has_sum (λ n, f (n + k)) a ↔ has_sum f (a + ∑ i in range k, f i) := @@ -922,7 +921,61 @@ begin exact (this (λ n, f n) hpos).nonneg_add_neg hneg, end -end subtype +lemma summable_int_of_summable_nat {f : ℤ → α} + (hp : summable (λ n:ℕ, f n)) (hn : summable (λ n:ℕ, f (-n))) : summable f := +(has_sum.nonneg_add_neg hp.has_sum $ summable.has_sum $ (summable_nat_add_iff 1).mpr hn).summable + +lemma has_sum.sum_nat_of_sum_int {α : Type*} [add_comm_monoid α] [topological_space α] + [has_continuous_add α] {a : α} {f : ℤ → α} (hf : has_sum f a) : + has_sum (λ n:ℕ, f n + f (-n)) (a + f 0) := +begin + apply (hf.add (has_sum_ite_eq (0 : ℤ) (f 0))).has_sum_of_sum_eq (λ u, _), + refine ⟨u.image int.nat_abs, λ v' hv', _⟩, + let u1 := v'.image (λ (x : ℕ), (x : ℤ)), + let u2 := v'.image (λ (x : ℕ), - (x : ℤ)), + have A : u ⊆ u1 ∪ u2, + { assume x hx, + simp only [mem_union, mem_image, exists_prop], + rcases le_total 0 x with h'x|h'x, + { left, + refine ⟨int.nat_abs x, hv' _, _⟩, + { simp only [mem_image, exists_prop], + exact ⟨x, hx, rfl⟩ }, + { simp only [h'x, int.coe_nat_abs, abs_eq_self] } }, + { right, + refine ⟨int.nat_abs x, hv' _, _⟩, + { simp only [mem_image, exists_prop], + exact ⟨x, hx, rfl⟩ }, + { simp only [abs_of_nonpos h'x, int.coe_nat_abs, neg_neg] } } }, + refine ⟨u1 ∪ u2, A, _⟩, + calc ∑ x in u1 ∪ u2, (f x + ite (x = 0) (f 0) 0) + = ∑ x in u1 ∪ u2, f x + ∑ x in u1 ∩ u2, f x : + begin + rw sum_add_distrib, + congr' 1, + refine (sum_subset_zero_on_sdiff inter_subset_union _ _).symm, + { assume x hx, + suffices : x ≠ 0, by simp only [this, if_false], + rintros rfl, + simpa only [mem_sdiff, mem_union, mem_image, neg_eq_zero, or_self, mem_inter, and_self, + and_not_self] using hx }, + { assume x hx, + simp only [mem_inter, mem_image, exists_prop] at hx, + have : x = 0, + { apply le_antisymm, + { rcases hx.2 with ⟨a, ha, rfl⟩, + simp only [right.neg_nonpos_iff, nat.cast_nonneg] }, + { rcases hx.1 with ⟨a, ha, rfl⟩, + simp only [nat.cast_nonneg] } }, + simp only [this, eq_self_iff_true, if_true] } + end + ... = ∑ x in u1, f x + ∑ x in u2, f x : sum_union_inter + ... = ∑ b in v', f b + ∑ b in v', f (-b) : + by simp only [sum_image, nat.cast_inj, imp_self, implies_true_iff, neg_inj] + ... = ∑ b in v', (f b + f (-b)) : sum_add_distrib.symm +end + +end nat end topological_group @@ -1352,16 +1405,6 @@ lemma tsum_comm [t1_space α] {f : β → γ → α} (h : summable (function.unc ∑' c b, f b c = ∑' b c, f b c := tsum_comm' h h.prod_factor h.prod_symm.prod_factor -lemma has_sum.sum_nat_of_sum_int [t2_space α] {f : ℤ → α} (hf : has_sum f a) : - has_sum (λ n:ℕ, f(n + 1) + f(-n.succ)) (a - f 0) := -begin - obtain ⟨b₁, h₁⟩ : summable (λ n : ℕ, f(n + 1)) := hf.summable.comp_injective (λ x₁ x₂, by simp), - obtain ⟨b₂, h₂⟩ : summable (λ n : ℕ, f(-n.succ)) := hf.summable.comp_injective (λ x₁ x₂, by simp), - convert h₁.add h₂, - rw hf.unique (h₁.pos_add_zero_add_neg h₂), - abel, -end - lemma tsum_subtype_add_tsum_subtype_compl [t2_space α] {f : β → α} (hf : summable f) (s : set β) : ∑' x : s, f x + ∑' x : sᶜ, f x = ∑' x, f x := ((hf.subtype s).has_sum.add_compl (hf.subtype {x | x ∉ s}).has_sum).unique hf.has_sum From 9f26ebf297c6a5ca26573a970411e606bb2ebe63 Mon Sep 17 00:00:00 2001 From: datokrat Date: Fri, 13 Jan 2023 22:32:37 +0000 Subject: [PATCH 0235/1291] feat(analysis/convex/intrinsic): Intrinsic interior and frontier (#18027) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Defines the intrinsic interior, closure and boundary of a set in a normed additive torsor (e.g., a real vector space or one of its nonempty affine subspaces). Results: - Simple lemmas about those definitions - `affine_isometry.image_intrinsic_interior`: The image of the intrinsic interior under an affine isometry is the relative interior of the image. - `set.nonempty.intrinsic_interior`: The intrinsic interior of a nonempty convex set is nonempty. Co-authored-by: Yaël Dillies Co-authored-by: Yaël Dillies --- src/analysis/convex/intrinsic.lean | 307 ++++++++++++++++++ .../affine_space/affine_equiv.lean | 6 + .../affine_space/affine_subspace.lean | 21 +- src/topology/homeomorph.lean | 3 + 4 files changed, 334 insertions(+), 3 deletions(-) create mode 100644 src/analysis/convex/intrinsic.lean diff --git a/src/analysis/convex/intrinsic.lean b/src/analysis/convex/intrinsic.lean new file mode 100644 index 0000000000000..035e2ea3e4686 --- /dev/null +++ b/src/analysis/convex/intrinsic.lean @@ -0,0 +1,307 @@ +/- +Copyright (c) 2023 Paul Reichert. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Paul Reichert, Yaël Dillies +-/ +import analysis.normed_space.add_torsor_bases +import analysis.normed_space.linear_isometry + +/-! +# Intrinsic frontier and interior + +This file defines the intrinsic frontier, interior and closure of a set in a normed additive torsor. +These are also known as relative frontier, interior, closure. + +The intrinsic frontier/interior/closure of a set `s` is the frontier/interior/closure of `s` +considered as a set in its affine span. + +The intrinsic interior is in general greater than the topological interior, the intrinsic frontier +in general less than the topological frontier, and the intrinsic closure in cases of interest the +same as the topological closure. + +## Definitions + +* `intrinsic_interior`: Intrinsic interior +* `intrinsic_frontier`: Intrinsic frontier +* `intrinsic_closure`: Intrinsic closure + +## Results + +The main results are: +* `affine_isometry.image_intrinsic_interior`/`affine_isometry.image_intrinsic_frontier`/ + `affine_isometry.image_intrinsic_closure`: Intrinsic interiors/frontiers/closures commute with + taking the image under an affine isometry. +* `set.nonempty.intrinsic_interior`: The intrinsic interior of a nonempty convex set is nonempty. + +## References + +* Chapter 8 of [Barry Simon, *Convexity*][simon2011] +* Chapter 1 of [Rolf Schneider, *Convex Bodies: The Brunn-Minkowski theory*][schneider2013]. + +## TODO + +* `is_closed s → is_extreme 𝕜 s (intrinsic_frontier 𝕜 s)` +* `x ∈ s → y ∈ intrinsic_interior 𝕜 s → open_segment 𝕜 x y ⊆ intrinsic_interior 𝕜 s` +-/ + +open affine_subspace set +open_locale pointwise + +variables {𝕜 V W Q P : Type*} + +section add_torsor +variables (𝕜) [ring 𝕜] [add_comm_group V] [module 𝕜 V] [topological_space P] [add_torsor V P] + {s t : set P} {x : P} +include V + +/-- The intrinsic interior of a set is its interior considered as a set in its affine span. -/ +def intrinsic_interior (s : set P) : set P := coe '' interior (coe ⁻¹' s : set $ affine_span 𝕜 s) + +/-- The intrinsic frontier of a set is its frontier considered as a set in its affine span. -/ +def intrinsic_frontier (s : set P) : set P := coe '' frontier (coe ⁻¹' s : set $ affine_span 𝕜 s) + +/-- The intrinsic closure of a set is its closure considered as a set in its affine span. -/ +def intrinsic_closure (s : set P) : set P := coe '' closure (coe ⁻¹' s : set $ affine_span 𝕜 s) + +variables {𝕜} + +@[simp] lemma mem_intrinsic_interior : + x ∈ intrinsic_interior 𝕜 s ↔ ∃ y, y ∈ interior (coe ⁻¹' s : set $ affine_span 𝕜 s) ∧ ↑y = x := +mem_image _ _ _ + +@[simp] lemma mem_intrinsic_frontier : + x ∈ intrinsic_frontier 𝕜 s ↔ ∃ y, y ∈ frontier (coe ⁻¹' s : set $ affine_span 𝕜 s) ∧ ↑y = x := +mem_image _ _ _ + +@[simp] lemma mem_intrinsic_closure : + x ∈ intrinsic_closure 𝕜 s ↔ ∃ y, y ∈ closure (coe ⁻¹' s : set $ affine_span 𝕜 s) ∧ ↑y = x := +mem_image _ _ _ + +lemma intrinsic_interior_subset : intrinsic_interior 𝕜 s ⊆ s := image_subset_iff.2 interior_subset + +lemma intrinsic_frontier_subset (hs : is_closed s) : intrinsic_frontier 𝕜 s ⊆ s := +image_subset_iff.2 (hs.preimage continuous_induced_dom).frontier_subset + +lemma intrinsic_frontier_subset_intrinsic_closure : + intrinsic_frontier 𝕜 s ⊆ intrinsic_closure 𝕜 s := +image_subset _ frontier_subset_closure + +lemma subset_intrinsic_closure : s ⊆ intrinsic_closure 𝕜 s := +λ x hx, ⟨⟨x, subset_affine_span _ _ hx⟩, subset_closure hx, rfl⟩ + +@[simp] lemma intrinsic_interior_empty : intrinsic_interior 𝕜 (∅ : set P) = ∅ := +by simp [intrinsic_interior] + +@[simp] lemma intrinsic_frontier_empty : intrinsic_frontier 𝕜 (∅ : set P) = ∅ := +by simp [intrinsic_frontier] + +@[simp] lemma intrinsic_closure_empty : intrinsic_closure 𝕜 (∅ : set P) = ∅ := +by simp [intrinsic_closure] + +@[simp] lemma intrinsic_closure_nonempty : (intrinsic_closure 𝕜 s).nonempty ↔ s.nonempty := +⟨by { simp_rw nonempty_iff_ne_empty, rintro h rfl, exact h intrinsic_closure_empty }, + nonempty.mono subset_intrinsic_closure⟩ + +alias intrinsic_closure_nonempty ↔ set.nonempty.of_intrinsic_closure set.nonempty.intrinsic_closure + +attribute [protected] set.nonempty.intrinsic_closure + +@[simp] lemma intrinsic_interior_singleton (x : P) : intrinsic_interior 𝕜 ({x} : set P) = {x} := +by simpa only [intrinsic_interior, preimage_coe_affine_span_singleton, interior_univ, image_univ, + subtype.range_coe] using coe_affine_span_singleton _ _ _ + +@[simp] lemma intrinsic_frontier_singleton (x : P) : intrinsic_frontier 𝕜 ({x} : set P) = ∅ := +by rw [intrinsic_frontier, preimage_coe_affine_span_singleton, frontier_univ, image_empty] + +@[simp] lemma intrinsic_closure_singleton (x : P) : intrinsic_closure 𝕜 ({x} : set P) = {x} := +by simpa only [intrinsic_closure, preimage_coe_affine_span_singleton, closure_univ, image_univ, + subtype.range_coe] using coe_affine_span_singleton _ _ _ + +/-! +Note that neither `intrinsic_interior` nor `intrinsic_frontier` is monotone. +-/ + +lemma intrinsic_closure_mono (h : s ⊆ t) : intrinsic_closure 𝕜 s ⊆ intrinsic_closure 𝕜 t := +begin + refine image_subset_iff.2 (λ x hx, ⟨set.inclusion (affine_span_mono _ h) x, + (continuous_inclusion _).closure_preimage_subset _ $ closure_mono _ hx, rfl⟩), + exact λ y hy, h hy, +end + +lemma interior_subset_intrinsic_interior : interior s ⊆ intrinsic_interior 𝕜 s := +λ x hx, ⟨⟨x, subset_affine_span _ _ $ interior_subset hx⟩, + preimage_interior_subset_interior_preimage continuous_subtype_coe hx, rfl⟩ + +lemma intrinsic_closure_subset_closure : intrinsic_closure 𝕜 s ⊆ closure s := +image_subset_iff.2 $ continuous_subtype_coe.closure_preimage_subset _ + +lemma intrinsic_frontier_subset_frontier : intrinsic_frontier 𝕜 s ⊆ frontier s := +image_subset_iff.2 $ continuous_subtype_coe.frontier_preimage_subset _ + +lemma intrinsic_closure_subset_affine_span : intrinsic_closure 𝕜 s ⊆ affine_span 𝕜 s := +(image_subset_range _ _).trans subtype.range_coe.subset + +@[simp] lemma intrinsic_closure_diff_intrinsic_frontier (s : set P) : + intrinsic_closure 𝕜 s \ intrinsic_frontier 𝕜 s = intrinsic_interior 𝕜 s := +(image_diff subtype.coe_injective _ _).symm.trans $ + by rw [closure_diff_frontier, intrinsic_interior] + +@[simp] lemma intrinsic_closure_diff_intrinsic_interior (s : set P) : + intrinsic_closure 𝕜 s \ intrinsic_interior 𝕜 s = intrinsic_frontier 𝕜 s := +(image_diff subtype.coe_injective _ _).symm + +@[simp] lemma intrinsic_interior_union_intrinsic_frontier (s : set P) : + intrinsic_interior 𝕜 s ∪ intrinsic_frontier 𝕜 s = intrinsic_closure 𝕜 s := +by simp [intrinsic_closure, intrinsic_interior, intrinsic_frontier, + closure_eq_interior_union_frontier, image_union] + +@[simp] lemma intrinsic_frontier_union_intrinsic_interior (s : set P) : + intrinsic_frontier 𝕜 s ∪ intrinsic_interior 𝕜 s = intrinsic_closure 𝕜 s := +by rw [union_comm, intrinsic_interior_union_intrinsic_frontier] + +lemma is_closed_intrinsic_closure (hs : is_closed (affine_span 𝕜 s : set P)) : + is_closed (intrinsic_closure 𝕜 s) := +(closed_embedding_subtype_coe hs).is_closed_map _ is_closed_closure + +lemma is_closed_intrinsic_frontier (hs : is_closed (affine_span 𝕜 s : set P)) : + is_closed (intrinsic_frontier 𝕜 s) := +(closed_embedding_subtype_coe hs).is_closed_map _ is_closed_frontier + +@[simp] lemma affine_span_intrinsic_closure (s : set P) : + affine_span 𝕜 (intrinsic_closure 𝕜 s) = affine_span 𝕜 s := +(affine_span_le.2 intrinsic_closure_subset_affine_span).antisymm $ + affine_span_mono _ subset_intrinsic_closure + +protected lemma is_closed.intrinsic_closure (hs : is_closed (coe ⁻¹' s : set $ affine_span 𝕜 s)) : + intrinsic_closure 𝕜 s = s := +begin + rw [intrinsic_closure, hs.closure_eq, image_preimage_eq_of_subset], + exact (subset_affine_span _ _).trans subtype.range_coe.superset, +end + +@[simp] lemma intrinsic_closure_idem (s : set P) : + intrinsic_closure 𝕜 (intrinsic_closure 𝕜 s) = intrinsic_closure 𝕜 s := +begin + refine is_closed.intrinsic_closure _, + set t := affine_span 𝕜 (intrinsic_closure 𝕜 s) with ht, + clear_value t, + obtain rfl := ht.trans (affine_span_intrinsic_closure _), + rw [intrinsic_closure, preimage_image_eq _ subtype.coe_injective], + exact is_closed_closure, +end + +end add_torsor + +namespace affine_isometry +variables [normed_field 𝕜] [seminormed_add_comm_group V] [seminormed_add_comm_group W] + [normed_space 𝕜 V] [normed_space 𝕜 W] [metric_space P] [pseudo_metric_space Q] + [normed_add_torsor V P] [normed_add_torsor W Q] +include V W + +local attribute [instance, nolint fails_quickly] affine_subspace.to_normed_add_torsor + affine_subspace.nonempty_map + +@[simp] lemma image_intrinsic_interior (φ : P →ᵃⁱ[𝕜] Q) (s : set P) : + intrinsic_interior 𝕜 (φ '' s) = φ '' intrinsic_interior 𝕜 s := +begin + obtain rfl | hs := s.eq_empty_or_nonempty, + { simp only [intrinsic_interior_empty, image_empty] }, + haveI : nonempty s := hs.to_subtype, + let f := ((affine_span 𝕜 s).isometry_equiv_map φ).to_homeomorph, + have : φ.to_affine_map ∘ coe ∘ f.symm = coe := funext isometry_equiv_map.apply_symm_apply, + rw [intrinsic_interior, intrinsic_interior, ←φ.coe_to_affine_map, ←map_span φ.to_affine_map s, + ←this, ←function.comp.assoc, image_comp, image_comp, f.symm.image_interior, f.image_symm, + ←preimage_comp, function.comp.assoc, f.symm_comp_self, affine_isometry.coe_to_affine_map, + function.comp.right_id, preimage_comp, φ.injective.preimage_image], +end + +@[simp] lemma image_intrinsic_frontier (φ : P →ᵃⁱ[𝕜] Q) (s : set P) : + intrinsic_frontier 𝕜 (φ '' s) = φ '' intrinsic_frontier 𝕜 s := +begin + obtain rfl | hs := s.eq_empty_or_nonempty, + { simp }, + haveI : nonempty s := hs.to_subtype, + let f := ((affine_span 𝕜 s).isometry_equiv_map φ).to_homeomorph, + have : φ.to_affine_map ∘ coe ∘ f.symm = coe := funext isometry_equiv_map.apply_symm_apply, + rw [intrinsic_frontier, intrinsic_frontier, ←φ.coe_to_affine_map, ←map_span φ.to_affine_map s, + ←this, ←function.comp.assoc, image_comp, image_comp, f.symm.image_frontier, f.image_symm, + ←preimage_comp, function.comp.assoc, f.symm_comp_self, affine_isometry.coe_to_affine_map, + function.comp.right_id, preimage_comp, φ.injective.preimage_image], +end + +@[simp] lemma image_intrinsic_closure (φ : P →ᵃⁱ[𝕜] Q) (s : set P) : + intrinsic_closure 𝕜 (φ '' s) = φ '' intrinsic_closure 𝕜 s := +begin + obtain rfl | hs := s.eq_empty_or_nonempty, + { simp }, + haveI : nonempty s := hs.to_subtype, + let f := ((affine_span 𝕜 s).isometry_equiv_map φ).to_homeomorph, + have : φ.to_affine_map ∘ coe ∘ f.symm = coe := funext isometry_equiv_map.apply_symm_apply, + rw [intrinsic_closure, intrinsic_closure, ←φ.coe_to_affine_map, ←map_span φ.to_affine_map s, + ←this, ←function.comp.assoc, image_comp, image_comp, f.symm.image_closure, f.image_symm, + ←preimage_comp, function.comp.assoc, f.symm_comp_self, affine_isometry.coe_to_affine_map, + function.comp.right_id, preimage_comp, φ.injective.preimage_image], +end + +end affine_isometry + +section normed_add_torsor +variables (𝕜) [nontrivially_normed_field 𝕜] [complete_space 𝕜] [normed_add_comm_group V] + [normed_space 𝕜 V] [finite_dimensional 𝕜 V] [metric_space P] [normed_add_torsor V P] (s : set P) +include V + +@[simp] lemma intrinsic_closure_eq_closure : intrinsic_closure 𝕜 s = closure s := +begin + ext x, + simp only [mem_closure_iff, mem_intrinsic_closure], + refine ⟨_, λ h, ⟨⟨x, _⟩, _, subtype.coe_mk _ _⟩⟩, + { rintro ⟨x, h, rfl⟩ t ht hx, + obtain ⟨z, hz₁, hz₂⟩ := h _ (continuous_induced_dom.is_open_preimage t ht) hx, + exact ⟨z, hz₁, hz₂⟩ }, + { by_contradiction hc, + obtain ⟨z, hz₁, hz₂⟩ := h _ (affine_span 𝕜 s).closed_of_finite_dimensional.is_open_compl hc, + exact hz₁ (subset_affine_span 𝕜 s hz₂) }, + { rintro _ ⟨t, ht, rfl⟩ hx, + obtain ⟨y, hyt, hys⟩ := h _ ht hx, + exact ⟨⟨_, subset_affine_span 𝕜 s hys⟩, hyt, hys⟩ } +end + +variables {𝕜} + +@[simp] lemma closure_diff_intrinsic_interior (s : set P) : + closure s \ intrinsic_interior 𝕜 s = intrinsic_frontier 𝕜 s := +intrinsic_closure_eq_closure 𝕜 s ▸ intrinsic_closure_diff_intrinsic_interior s + +@[simp] lemma closure_diff_intrinsic_frontier (s : set P) : + closure s \ intrinsic_frontier 𝕜 s = intrinsic_interior 𝕜 s := +intrinsic_closure_eq_closure 𝕜 s ▸ intrinsic_closure_diff_intrinsic_frontier s + +end normed_add_torsor + +private lemma aux {α β : Type*} [topological_space α] [topological_space β] (φ : α ≃ₜ β) + (s : set β) : + (interior s).nonempty ↔ (interior (φ ⁻¹' s)).nonempty := +by rw [←φ.image_symm, ←φ.symm.image_interior, nonempty_image_iff] + +variables [normed_add_comm_group V] [normed_space ℝ V] [finite_dimensional ℝ V] {s : set V} + +/-- The intrinsic interior of a nonempty convex set is nonempty. -/ +protected lemma set.nonempty.intrinsic_interior (hscv : convex ℝ s) (hsne : s.nonempty) : + (intrinsic_interior ℝ s).nonempty := +begin + haveI := hsne.coe_sort, + obtain ⟨p, hp⟩ := hsne, + let p' : affine_span ℝ s := ⟨p, subset_affine_span _ _ hp⟩, + rw [intrinsic_interior, nonempty_image_iff, + aux (affine_isometry_equiv.const_vsub ℝ p').symm.to_homeomorph, + convex.interior_nonempty_iff_affine_span_eq_top, affine_isometry_equiv.coe_to_homeomorph, + ←affine_isometry_equiv.coe_to_affine_equiv, ←comap_span, affine_span_coe_preimage_eq_top, + comap_top], + exact hscv.affine_preimage ((affine_span ℝ s).subtype.comp + (affine_isometry_equiv.const_vsub ℝ p').symm.to_affine_equiv.to_affine_map), +end + +lemma intrinsic_interior_nonempty (hs : convex ℝ s) : + (intrinsic_interior ℝ s).nonempty ↔ s.nonempty := +⟨by { simp_rw nonempty_iff_ne_empty, rintro h rfl, exact h intrinsic_interior_empty }, + set.nonempty.intrinsic_interior hs⟩ diff --git a/src/linear_algebra/affine_space/affine_equiv.lean b/src/linear_algebra/affine_space/affine_equiv.lean index fb25afd2fccdc..72b42e85a3e29 100644 --- a/src/linear_algebra/affine_space/affine_equiv.lean +++ b/src/linear_algebra/affine_space/affine_equiv.lean @@ -184,6 +184,12 @@ e.to_equiv.apply_eq_iff_eq_symm_apply @[simp] lemma apply_eq_iff_eq (e : P₁ ≃ᵃ[k] P₂) {p₁ p₂ : P₁} : e p₁ = e p₂ ↔ p₁ = p₂ := e.to_equiv.apply_eq_iff_eq +@[simp] lemma image_symm (f : P₁ ≃ᵃ[k] P₂) (s : set P₂) : f.symm '' s = f ⁻¹' s := +f.symm.to_equiv.image_eq_preimage _ + +@[simp] lemma preimage_symm (f : P₁ ≃ᵃ[k] P₂) (s : set P₁) : f.symm ⁻¹' s = f '' s := +(f.symm.image_symm _).symm + variables (k P₁) omit V₂ diff --git a/src/linear_algebra/affine_space/affine_subspace.lean b/src/linear_algebra/affine_space/affine_subspace.lean index 83e3fddab70b9..d7e5ebcb29fe2 100644 --- a/src/linear_algebra/affine_space/affine_subspace.lean +++ b/src/linear_algebra/affine_space/affine_subspace.lean @@ -675,7 +675,7 @@ lemma _root_.affine_span_le {s : set P} {Q : affine_subspace k P} : affine_span k s ≤ Q ↔ s ⊆ (Q : set P) := (affine_subspace.gi k V P).gc _ _ -variables (k V) {P} +variables (k V) {P} {p₁ p₂ : P} /-- The affine span of a single point, coerced to a set, contains just that point. -/ @@ -689,10 +689,13 @@ end /-- A point is in the affine span of a single point if and only if they are equal. -/ -@[simp] lemma mem_affine_span_singleton (p1 p2 : P) : - p1 ∈ affine_span k ({p2} : set P) ↔ p1 = p2 := +@[simp] lemma mem_affine_span_singleton : p₁ ∈ affine_span k ({p₂} : set P) ↔ p₁ = p₂ := by simp [←mem_coe] +@[simp] lemma preimage_coe_affine_span_singleton (x : P) : + (coe : affine_span k ({x} : set P) → P) ⁻¹' {x} = univ := +eq_univ_of_forall $ λ y, (affine_subspace.mem_affine_span_singleton _ _).1 y.2 + /-- The span of a union of sets is the sup of their spans. -/ lemma span_union (s t : set P) : affine_span k (s ∪ t) = affine_span k s ⊔ affine_span k t := (affine_subspace.gi k V P).gc.l_sup @@ -1598,6 +1601,18 @@ lemma comap_supr {ι : Sort*} (f : P₁ →ᵃ[k] P₂) (s : ι → affine_subsp (infi s).comap f = ⨅ i, (s i).comap f := (gc_map_comap f).u_infi +@[simp] lemma comap_symm (e : P₁ ≃ᵃ[k] P₂) (s : affine_subspace k P₁) : + s.comap (e.symm : P₂ →ᵃ[k] P₁) = s.map e := +coe_injective $ e.preimage_symm _ + +@[simp] lemma map_symm (e : P₁ ≃ᵃ[k] P₂) (s : affine_subspace k P₂) : + s.map (e.symm : P₂ →ᵃ[k] P₁) = s.comap e := +coe_injective $ e.image_symm _ + +lemma comap_span (f : P₁ ≃ᵃ[k] P₂) (s : set P₂) : + (affine_span k s).comap (f : P₁ →ᵃ[k] P₂) = affine_span k (f ⁻¹' s) := +by rw [←map_symm, map_span, affine_equiv.coe_coe, f.image_symm] + end affine_subspace end map_comap diff --git a/src/topology/homeomorph.lean b/src/topology/homeomorph.lean index edadf75d322b0..3d0ef7172be6c 100644 --- a/src/topology/homeomorph.lean +++ b/src/topology/homeomorph.lean @@ -249,6 +249,9 @@ by rw [← preimage_symm, preimage_interior] lemma preimage_frontier (h : α ≃ₜ β) (s : set β) : h ⁻¹' (frontier s) = frontier (h ⁻¹' s) := h.is_open_map.preimage_frontier_eq_frontier_preimage h.continuous _ +lemma image_frontier (h : α ≃ₜ β) (s : set α) : h '' frontier s = frontier (h '' s) := +by rw [←preimage_symm, preimage_frontier] + @[to_additive] lemma _root_.has_compact_mul_support.comp_homeomorph {M} [has_one M] {f : β → M} (hf : has_compact_mul_support f) (φ : α ≃ₜ β) : has_compact_mul_support (f ∘ φ) := From c3350dbc6017f74483b83aa5760dec3bae19fcba Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Fri, 13 Jan 2023 22:32:38 +0000 Subject: [PATCH 0236/1291] feat(ring_theory): equality of ideals from locally at each maximal ideal (#18142) This PR shows that an ideal is equal to another if all localizations at each maximal ideal are equal, by showing an ideal is included in another if this holds locally at each maximal ideal. This generalizes `ideal_eq_zero_of_localization`. The proof is inspired somewhat by the one for `maximal_spectrum.infi_localization_eq_bot`, although I couldn't find out a neat way to prove it from the new results since it talks about slightly different maps. It would also require readjusting the imports. Zulip thread: https://leanprover.zulipchat.com/#narrow/stream/217875-Is-there-code-for-X.3F/topic/Equality.20of.20ideals.20from.20local.20equality Co-authored-by: Anne Baanen --- .../group_action/sub_mul_action.lean | 7 ++ src/ring_theory/ideal/operations.lean | 11 +++ src/ring_theory/local_properties.lean | 83 ++++++++++++------- 3 files changed, 71 insertions(+), 30 deletions(-) diff --git a/src/group_theory/group_action/sub_mul_action.lean b/src/group_theory/group_action/sub_mul_action.lean index 30e466ad411cd..96abf240857d2 100644 --- a/src/group_theory/group_action/sub_mul_action.lean +++ b/src/group_theory/group_action/sub_mul_action.lean @@ -70,6 +70,13 @@ lemma mk_smul_mk (r : R) (x : M) (hx : x ∈ s) : @[to_additive] lemma smul_def (r : R) (x : s) : r • x = ⟨r • x, smul_mem r x.2⟩ := rfl +omit hS + +@[simp] lemma forall_smul_mem_iff {R M S : Type*} [monoid R] [mul_action R M] + [set_like S M] [smul_mem_class S R M] {N : S} {x : M} : + (∀ (a : R), a • x ∈ N) ↔ x ∈ N := +⟨λ h, by simpa using h 1, λ h a, smul_mem_class.smul_mem a h⟩ + end set_like /-- A sub_mul_action is a set which is closed under scalar multiplication. -/ diff --git a/src/ring_theory/ideal/operations.lean b/src/ring_theory/ideal/operations.lean index 35bd9658d727f..21e13f9e69593 100644 --- a/src/ring_theory/ideal/operations.lean +++ b/src/ring_theory/ideal/operations.lean @@ -319,6 +319,17 @@ le_antisymm (le_infi $ λ i, le_infi $ λ j, colon_mono (infi_le _ _) (le_supr _ map_le_iff_le_comap.2 $ mem_colon'.1 $ have _ := ((mem_infi _).1 H i), have _ := ((mem_infi _).1 this j), this) +@[simp] lemma mem_colon_singleton {N : submodule R M} {x : M} {r : R} : + r ∈ N.colon (submodule.span R {x}) ↔ r • x ∈ N := +calc r ∈ N.colon (submodule.span R {x}) ↔ ∀ (a : R), r • (a • x) ∈ N : + by simp [submodule.mem_colon, submodule.mem_span_singleton] + ... ↔ r • x ∈ N : + by { simp_rw [smul_comm r]; exact set_like.forall_smul_mem_iff } + +@[simp] lemma _root_.ideal.mem_colon_singleton {I : ideal R} {x r : R} : + r ∈ I.colon (ideal.span {x}) ↔ r * x ∈ I := +by simp [← ideal.submodule_span_eq, submodule.mem_colon_singleton, smul_eq_mul] + end comm_ring end submodule diff --git a/src/ring_theory/local_properties.lean b/src/ring_theory/local_properties.lean index b862b968f5e37..1e60de6874cf7 100644 --- a/src/ring_theory/local_properties.lean +++ b/src/ring_theory/local_properties.lean @@ -29,7 +29,7 @@ In this file, we provide the proofs of various local properties. The following properties are covered: * The triviality of an ideal or an element: - `ideal_eq_zero_of_localization`, `eq_zero_of_localization` + `ideal_eq_bot_of_localization`, `eq_zero_of_localization` * `is_reduced` : `localization_is_reduced`, `is_reduced_of_localization_maximal`. * `finite`: `localization_finite`, `finite_of_localization_span` * `finite_type`: `localization_finite_type`, `finite_type_of_localization_span` @@ -221,43 +221,66 @@ end properties section ideal --- This proof should work for all modules, but we do not know how to localize a module yet. -/-- An ideal is trivial if its localization at every maximal ideal is trivial. -/ -lemma ideal_eq_zero_of_localization (I : ideal R) - (h : ∀ (J : ideal R) (hJ : J.is_maximal), - by exactI is_localization.coe_submodule (localization.at_prime J) I = 0) : I = 0 := +open_locale non_zero_divisors + +/-- Let `I J : ideal R`. If the localization of `I` at each maximal ideal `P` is included in +the localization of `J` at `P`, then `I ≤ J`. -/ +lemma ideal.le_of_localization_maximal {I J : ideal R} + (h : ∀ (P : ideal R) (hP : P.is_maximal), + ideal.map (algebra_map R (by exactI localization.at_prime P)) I ≤ + ideal.map (algebra_map R (by exactI localization.at_prime P)) J) : + I ≤ J := begin - by_contradiction hI, change I ≠ ⊥ at hI, - obtain ⟨x, hx, hx'⟩ := set_like.exists_of_lt hI.bot_lt, - rw [submodule.mem_bot] at hx', - have H : (ideal.span ({x} : set R)).annihilator ≠ ⊤, - { rw [ne.def, submodule.annihilator_eq_top_iff], - by_contra, - apply hx', - rw [← set.mem_singleton_iff, ← @submodule.bot_coe R, ← h], - exact ideal.subset_span (set.mem_singleton x) }, - obtain ⟨p, hp₁, hp₂⟩ := ideal.exists_le_maximal _ H, - resetI, - specialize h p hp₁, - have : algebra_map R (localization.at_prime p) x = 0, - { rw ← set.mem_singleton_iff, - change algebra_map R (localization.at_prime p) x ∈ (0 : submodule R (localization.at_prime p)), - rw ← h, - exact submodule.mem_map_of_mem hx }, - rw is_localization.map_eq_zero_iff p.prime_compl at this, - obtain ⟨m, hm⟩ := this, - apply m.prop, - refine hp₂ _, - erw submodule.mem_annihilator_span_singleton, - rwa mul_comm at hm, + intros x hx, + suffices : J.colon (ideal.span {x}) = ⊤, + { simpa using submodule.mem_colon.mp + (show (1 : R) ∈ J.colon (ideal.span {x}), from this.symm ▸ submodule.mem_top) + x (ideal.mem_span_singleton_self x) }, + refine not.imp_symm (J.colon (ideal.span {x})).exists_le_maximal _, + push_neg, + introsI P hP le, + obtain ⟨⟨⟨a, ha⟩, ⟨s, hs⟩⟩, eq⟩ := (is_localization.mem_map_algebra_map_iff P.prime_compl _).mp + (h P hP (ideal.mem_map_of_mem _ hx)), + rw [← _root_.map_mul, ← sub_eq_zero, ← map_sub] at eq, + obtain ⟨⟨m, hm⟩, eq⟩ := (is_localization.map_eq_zero_iff P.prime_compl _ _).mp eq, + refine hs ((hP.is_prime.mem_or_mem (le (ideal.mem_colon_singleton.mpr _))).resolve_right hm), + simp only [subtype.coe_mk, sub_mul, sub_eq_zero, mul_assoc] at eq, + simpa only [eq, mul_comm] using J.mul_mem_right m ha end +/-- Let `I J : ideal R`. If the localization of `I` at each maximal ideal `P` is equal to +the localization of `J` at `P`, then `I = J`. -/ +theorem ideal.eq_of_localization_maximal {I J : ideal R} + (h : ∀ (P : ideal R) (hP : P.is_maximal), + ideal.map (algebra_map R (by exactI localization.at_prime P)) I = + ideal.map (algebra_map R (by exactI localization.at_prime P)) J) : + I = J := +le_antisymm + (ideal.le_of_localization_maximal (λ P hP, (h P hP).le)) + (ideal.le_of_localization_maximal (λ P hP, (h P hP).ge)) + +/-- An ideal is trivial if its localization at every maximal ideal is trivial. -/ +lemma ideal_eq_bot_of_localization' (I : ideal R) + (h : ∀ (J : ideal R) (hJ : J.is_maximal), + ideal.map (algebra_map R (by exactI (localization.at_prime J))) I = ⊥) : I = ⊥ := +ideal.eq_of_localization_maximal (λ P hP, (by simpa using h P hP)) + +-- TODO: This proof should work for all modules, once we have enough material on submodules of +-- localized modules. +/-- An ideal is trivial if its localization at every maximal ideal is trivial. -/ +lemma ideal_eq_bot_of_localization (I : ideal R) + (h : ∀ (J : ideal R) (hJ : J.is_maximal), + by exactI is_localization.coe_submodule (localization.at_prime J) I = ⊥) : I = ⊥ := +ideal_eq_bot_of_localization' _ (λ P hP, (ideal.map_eq_bot_iff_le_ker _).mpr (λ x hx, + by { rw [ring_hom.mem_ker, ← submodule.mem_bot R, ← h P hP, is_localization.mem_coe_submodule], + exact ⟨x, hx, rfl⟩ })) + lemma eq_zero_of_localization (r : R) (h : ∀ (J : ideal R) (hJ : J.is_maximal), by exactI algebra_map R (localization.at_prime J) r = 0) : r = 0 := begin rw ← ideal.span_singleton_eq_bot, - apply ideal_eq_zero_of_localization, + apply ideal_eq_bot_of_localization, intros J hJ, delta is_localization.coe_submodule, erw [submodule.map_span, submodule.span_eq_bot], From 3310acfa9787aa171db6d4cba3945f6f275fe9f2 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Fri, 13 Jan 2023 22:32:39 +0000 Subject: [PATCH 0237/1291] feat(order/bounds/basic): add `is_greatest_univ_iff` etc (#18162) Also add `ae_measurable_const'` and golf 2 proofs. --- .../constructions/borel_space.lean | 39 +++++-------------- src/measure_theory/measure/ae_measurable.lean | 16 +++++--- src/order/bounds/basic.lean | 29 +++++++------- 3 files changed, 36 insertions(+), 48 deletions(-) diff --git a/src/measure_theory/constructions/borel_space.lean b/src/measure_theory/constructions/borel_space.lean index e97b2770a683a..ae3b12da0eabf 100644 --- a/src/measure_theory/constructions/borel_space.lean +++ b/src/measure_theory/constructions/borel_space.lean @@ -1049,51 +1049,32 @@ begin exact measurable_set.Union (λ i, hf i (is_open_gt' _).measurable_set) end -private lemma ae_measurable.is_glb_of_nonempty {ι} (hι : nonempty ι) - {μ : measure δ} [countable ι] {f : ι → δ → α} {g : δ → α} +lemma ae_measurable.is_glb {ι} {μ : measure δ} [countable ι] {f : ι → δ → α} {g : δ → α} (hf : ∀ i, ae_measurable (f i) μ) (hg : ∀ᵐ b ∂μ, is_glb {a | ∃ i, f i b = a} (g b)) : ae_measurable g μ := begin + nontriviality α, + haveI hα : nonempty α := infer_instance, + casesI is_empty_or_nonempty ι with hι hι, + { simp only [is_empty.exists_iff, set_of_false, is_glb_empty_iff] at hg, + exact ae_measurable_const' (hg.mono $ λ a ha, hg.mono $ λ b hb, (hb _).antisymm (ha _)) }, let p : δ → (ι → α) → Prop := λ x f', is_glb {a | ∃ i, f' i = a} (g x), - let g_seq := λ x, ite (x ∈ ae_seq_set hf p) (g x) (⟨g x⟩ : nonempty α).some, + let g_seq := (ae_seq_set hf p).piecewise g (λ _, hα.some), have hg_seq : ∀ b, is_glb {a | ∃ i, ae_seq hf p i b = a} (g_seq b), { intro b, - haveI hα : nonempty α := nonempty.map g ⟨b⟩, - simp only [ae_seq, g_seq], + simp only [ae_seq, g_seq, set.piecewise], split_ifs, { have h_set_eq : {a : α | ∃ (i : ι), (hf i).mk (f i) b = a} = {a : α | ∃ (i : ι), f i b = a}, { ext x, simp_rw [set.mem_set_of_eq, ae_seq.mk_eq_fun_of_mem_ae_seq_set hf h], }, rw h_set_eq, exact ae_seq.fun_prop_of_mem_ae_seq_set hf h, }, - { have h_singleton : {a : α | ∃ (i : ι), hα.some = a} = {hα.some}, - { ext1 x, - exact ⟨λ hx, hx.some_spec.symm, λ hx, ⟨hι.some, hx.symm⟩⟩, }, - rw h_singleton, - exact is_glb_singleton, }, }, + { exact is_least.is_glb ⟨(@exists_const (hα.some = hα.some) ι _).2 rfl, λ x ⟨i, hi⟩, hi.le⟩ } }, refine ⟨g_seq, measurable.is_glb (ae_seq.measurable hf p) hg_seq, _⟩, - exact (ite_ae_eq_of_measure_compl_zero g (λ x, (⟨g x⟩ : nonempty α).some) (ae_seq_set hf p) + exact (ite_ae_eq_of_measure_compl_zero g (λ x, hα.some) (ae_seq_set hf p) (ae_seq.measure_compl_ae_seq_set_eq_zero hf hg)).symm, end -lemma ae_measurable.is_glb {ι} {μ : measure δ} [countable ι] {f : ι → δ → α} {g : δ → α} - (hf : ∀ i, ae_measurable (f i) μ) (hg : ∀ᵐ b ∂μ, is_glb {a | ∃ i, f i b = a} (g b)) : - ae_measurable g μ := -begin - by_cases hμ : μ = 0, { rw hμ, exact ae_measurable_zero_measure }, - haveI : μ.ae.ne_bot, { simpa [ne_bot_iff] }, - by_cases hι : nonempty ι, { exact ae_measurable.is_glb_of_nonempty hι hf hg, }, - suffices : ∃ x, g =ᵐ[μ] λ y, g x, - by { exact ⟨(λ y, g this.some), measurable_const, this.some_spec⟩, }, - have h_empty : ∀ x, {a : α | ∃ (i : ι), f i x = a} = ∅, - { intro x, - ext1 y, - rw [set.mem_set_of_eq, set.mem_empty_iff_false, iff_false], - exact λ hi, hι (nonempty_of_exists hi), }, - simp_rw h_empty at hg, - exact ⟨hg.exists.some, hg.mono (λ y hy, is_glb.unique hy hg.exists.some_spec)⟩, -end - protected lemma monotone.measurable [linear_order β] [order_closed_topology β] {f : β → α} (hf : monotone f) : measurable f := suffices h : ∀ x, ord_connected (f ⁻¹' Ioi x), diff --git a/src/measure_theory/measure/ae_measurable.lean b/src/measure_theory/measure/ae_measurable.lean index c947c59a067b5..6860189ef0373 100644 --- a/src/measure_theory/measure/ae_measurable.lean +++ b/src/measure_theory/measure/ae_measurable.lean @@ -208,6 +208,15 @@ let ⟨g, hgm, hg⟩ := h in hgm.null_measurable.congr hg.symm end ae_measurable +lemma ae_measurable_const' (h : ∀ᵐ x y ∂μ, f x = f y) : ae_measurable f μ := +begin + rcases eq_or_ne μ 0 with rfl | hμ, + { exact ae_measurable_zero_measure }, + { haveI := ae_ne_bot.2 hμ, + rcases h.exists with ⟨x, hx⟩, + exact ⟨const α (f x), measurable_const, eventually_eq.symm hx⟩ } +end + lemma ae_measurable_uIoc_iff [linear_order α] {f : α → β} {a b : α} : (ae_measurable f $ μ.restrict $ Ι a b) ↔ (ae_measurable f $ μ.restrict $ Ioc a b) ∧ (ae_measurable f $ μ.restrict $ Ioc b a) := @@ -274,18 +283,15 @@ lemma ae_measurable_Ioi_of_forall_Ioc {β} {mβ : measurable_space β} ae_measurable g (μ.restrict (Ioi x)) := begin haveI : nonempty α := ⟨x⟩, - haveI : (at_top : filter α).ne_bot := at_top_ne_bot, obtain ⟨u, hu_tendsto⟩ := exists_seq_tendsto (at_top : filter α), have Ioi_eq_Union : Ioi x = ⋃ n : ℕ, Ioc x (u n), { rw Union_Ioc_eq_Ioi_self_iff.mpr _, - rw tendsto_at_top_at_top at hu_tendsto, - exact λ y _, ⟨(hu_tendsto y).some, (hu_tendsto y).some_spec (hu_tendsto y).some le_rfl⟩, }, + exact λ y _, (hu_tendsto.eventually (eventually_ge_at_top y)).exists }, rw [Ioi_eq_Union, ae_measurable_Union_iff], intros n, cases lt_or_le x (u n), { exact g_meas (u n) h, }, - { rw Ioc_eq_empty (not_lt.mpr h), - simp only [measure.restrict_empty], + { rw [Ioc_eq_empty (not_lt.mpr h), measure.restrict_empty], exact ae_measurable_zero_measure, }, end diff --git a/src/order/bounds/basic.lean b/src/order/bounds/basic.lean index e688e36d34499..6cfd3c71db1d3 100644 --- a/src/order/bounds/basic.lean +++ b/src/order/bounds/basic.lean @@ -571,25 +571,25 @@ by simp only [Ici_inter_Iic.symm, subset_inter_iff, bdd_below_iff_subset_Ici, #### Univ -/ -lemma is_greatest_univ [preorder γ] [order_top γ] : is_greatest (univ : set γ) ⊤ := -⟨mem_univ _, λ x hx, le_top⟩ +@[simp] lemma is_greatest_univ_iff : is_greatest univ a ↔ is_top a := +by simp [is_greatest, mem_upper_bounds, is_top] + +lemma is_greatest_univ [order_top α] : is_greatest (univ : set α) ⊤ := +is_greatest_univ_iff.2 is_top_top @[simp] lemma order_top.upper_bounds_univ [partial_order γ] [order_top γ] : upper_bounds (univ : set γ) = {⊤} := by rw [is_greatest_univ.upper_bounds_eq, Ici_top] -lemma is_lub_univ [preorder γ] [order_top γ] : is_lub (univ : set γ) ⊤ := -is_greatest_univ.is_lub +lemma is_lub_univ [order_top α] : is_lub (univ : set α) ⊤ := 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 γᵒᵈ _ _ -lemma is_least_univ [preorder γ] [order_bot γ] : is_least (univ : set γ) ⊥ := -@is_greatest_univ γᵒᵈ _ _ - -lemma is_glb_univ [preorder γ] [order_bot γ] : is_glb (univ : set γ) ⊥ := -is_least_univ.is_glb +@[simp] lemma is_least_univ_iff : is_least univ a ↔ is_bot a := @is_greatest_univ_iff αᵒᵈ _ _ +lemma is_least_univ [order_bot α] : is_least (univ : set α) ⊥ := @is_greatest_univ αᵒᵈ _ _ +lemma is_glb_univ [order_bot α] : is_glb (univ : set α) ⊥ := is_least_univ.is_glb @[simp] lemma no_max_order.upper_bounds_univ [no_max_order α] : upper_bounds (univ : set α) = ∅ := eq_empty_of_subset_empty $ λ b hb, let ⟨x, hx⟩ := exists_gt b in @@ -619,10 +619,11 @@ by simp only [bdd_above, upper_bounds_empty, univ_nonempty] @[simp] lemma bdd_below_empty [nonempty α] : bdd_below (∅ : set α) := 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] +@[simp] lemma is_glb_empty_iff : is_glb ∅ a ↔ is_top a := by simp [is_glb] +@[simp] lemma is_lub_empty_iff : is_lub ∅ a ↔ is_bot a := @is_glb_empty_iff αᵒᵈ _ _ -lemma is_lub_empty [preorder γ] [order_bot γ] : is_lub ∅ (⊥:γ) := @is_glb_empty γᵒᵈ _ _ +lemma is_glb_empty [order_top α] : is_glb ∅ (⊤:α) := is_glb_empty_iff.2 is_top_top +lemma is_lub_empty [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 @@ -683,11 +684,11 @@ by rw [insert_eq, upper_bounds_union, upper_bounds_singleton] by rw [insert_eq, lower_bounds_union, lower_bounds_singleton] /-- When there is a global maximum, every set is bounded above. -/ -@[simp] protected lemma order_top.bdd_above [preorder γ] [order_top γ] (s : set γ) : bdd_above s := +@[simp] protected lemma order_top.bdd_above [order_top α] (s : set α) : bdd_above s := ⟨⊤, λ a ha, order_top.le_top a⟩ /-- When there is a global minimum, every set is bounded below. -/ -@[simp] protected lemma order_bot.bdd_below [preorder γ] [order_bot γ] (s : set γ) : bdd_below s := +@[simp] protected lemma order_bot.bdd_below [order_bot α] (s : set α) : bdd_below s := ⟨⊥, λ a ha, order_bot.bot_le a⟩ /-! From 80c43012d26f63026d362c3aba28f3c3bafb07e6 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sat, 14 Jan 2023 03:06:06 +0000 Subject: [PATCH 0238/1291] feat(data/part): add 2 lemmas (#18153) --- src/data/part.lean | 8 ++++++++ 1 file changed, 8 insertions(+) diff --git a/src/data/part.lean b/src/data/part.lean index a771023f74a83..bd31d73beec2b 100644 --- a/src/data/part.lean +++ b/src/data/part.lean @@ -70,6 +70,14 @@ variables {α : Type*} {β : Type*} {γ : Type*} def to_option (o : part α) [decidable o.dom] : option α := if h : dom o then some (o.get h) else none +@[simp] lemma to_option_is_some (o : part α) [decidable o.dom] : + o.to_option.is_some ↔ o.dom := +by by_cases o.dom; simp [h, part.to_option] + +@[simp] lemma to_option_is_none (o : part α) [decidable o.dom] : + o.to_option.is_none ↔ ¬o.dom := +by by_cases o.dom; simp [h, part.to_option] + /-- `part` extensionality -/ theorem ext' : ∀ {o p : part α} (H1 : o.dom ↔ p.dom) From 00638177efd1b2534fc5269363ebf42a7871df9a Mon Sep 17 00:00:00 2001 From: Chris Hughes Date: Sat, 14 Jan 2023 03:06:08 +0000 Subject: [PATCH 0239/1291] chore(data/multiset/sort): make multiset repr a meta instance (#18163) Co-authored-by: Chris Hughes <33847686+ChrisHughes24@users.noreply.github.com> Co-authored-by: Eric Wieser --- src/data/finset/sort.lean | 2 +- src/data/list/cycle.lean | 14 +++++++------- src/data/multiset/sort.lean | 6 +++--- src/data/pnat/factors.lean | 5 ++++- src/group_theory/perm/cycle/concrete.lean | 2 +- test/cycle.lean | 4 ++-- 6 files changed, 18 insertions(+), 15 deletions(-) diff --git a/src/data/finset/sort.lean b/src/data/finset/sort.lean index 1ced1d74ff4f2..c293a1d344bf9 100644 --- a/src/data/finset/sort.lean +++ b/src/data/finset/sort.lean @@ -211,6 +211,6 @@ by simp only [order_emb_of_card_le, rel_embedding.coe_trans, finset.order_emb_of end sort_linear_order -instance [has_repr α] : has_repr (finset α) := ⟨λ s, repr s.1⟩ +meta instance [has_repr α] : has_repr (finset α) := ⟨λ s, repr s.1⟩ end finset diff --git a/src/data/list/cycle.lean b/src/data/list/cycle.lean index 60dc82c3a7e7b..8cfba07a9c6a6 100644 --- a/src/data/list/cycle.lean +++ b/src/data/list/cycle.lean @@ -17,8 +17,8 @@ Based on this, we define the quotient of lists by the rotation relation, called We also define a representation of concrete cycles, available when viewing them in a goal state or via `#eval`, when over representatble types. For example, the cycle `(2 1 4 3)` will be shown -as `c[1, 4, 3, 2]`. The representation of the cycle sorts the elements by the string value of the -underlying element. This representation also supports cycles that can contain duplicates. +as `c[2, 1, 4, 3]`. Two equal cycles may be printed differently if their internal representation +is different. -/ @@ -728,12 +728,12 @@ end decidable /-- We define a representation of concrete cycles, available when viewing them in a goal state or -via `#eval`, when over representatble types. For example, the cycle `(2 1 4 3)` will be shown -as `c[1, 4, 3, 2]`. The representation of the cycle sorts the elements by the string value of the -underlying element. This representation also supports cycles that can contain duplicates. +via `#eval`, when over representable types. For example, the cycle `(2 1 4 3)` will be shown +as `c[2, 1, 4, 3]`. Two equal cycles may be printed differently if their internal representation +is different. -/ -instance [has_repr α] : has_repr (cycle α) := -⟨λ s, "c[" ++ string.intercalate ", " ((s.map repr).lists.sort (≤)).head ++ "]"⟩ +meta instance [has_repr α] : has_repr (cycle α) := +⟨λ s, "c[" ++ string.intercalate ", " ((s.map repr).lists.unquot).head ++ "]"⟩ /-- `chain R s` means that `R` holds between adjacent elements of `s`. diff --git a/src/data/multiset/sort.lean b/src/data/multiset/sort.lean index df8ecbe0d9e4b..648d975765165 100644 --- a/src/data/multiset/sort.lean +++ b/src/data/multiset/sort.lean @@ -5,7 +5,6 @@ Authors: Mario Carneiro -/ import data.list.sort import data.multiset.basic -import data.string.basic /-! # Construct a sorted list from a multiset. @@ -50,7 +49,8 @@ list.merge_sort_singleton r a end sort -instance [has_repr α] : has_repr (multiset α) := -⟨λ s, "{" ++ string.intercalate ", " ((s.map repr).sort (≤)) ++ "}"⟩ +-- TODO: use a sort order if available, gh-18166 +meta instance [has_repr α] : has_repr (multiset α) := +⟨λ s, "{" ++ string.intercalate ", " (s.unquot.map repr) ++ "}"⟩ end multiset diff --git a/src/data/pnat/factors.lean b/src/data/pnat/factors.lean index 40fd6f53fc58c..f49517117765b 100644 --- a/src/data/pnat/factors.lean +++ b/src/data/pnat/factors.lean @@ -24,12 +24,15 @@ the multiplicity of `p` in this factors multiset being the p-adic valuation of ` /-- The type of multisets of prime numbers. Unique factorization gives an equivalence between this set and ℕ+, as we will formalize below. -/ - @[derive [inhabited, has_repr, canonically_ordered_add_monoid, distrib_lattice, +@[derive [inhabited, canonically_ordered_add_monoid, distrib_lattice, semilattice_sup, order_bot, has_sub, has_ordered_sub]] def prime_multiset := multiset nat.primes namespace prime_multiset +-- `@[derive]` doesn't work for `meta` instances +meta instance : has_repr prime_multiset := by delta prime_multiset; apply_instance + /-- The multiset consisting of a single prime -/ def of_prime (p : nat.primes) : prime_multiset := ({p} : multiset nat.primes) diff --git a/src/group_theory/perm/cycle/concrete.lean b/src/group_theory/perm/cycle/concrete.lean index 03a7eb22a2180..d52b8721d9ebc 100644 --- a/src/group_theory/perm/cycle/concrete.lean +++ b/src/group_theory/perm/cycle/concrete.lean @@ -537,7 +537,7 @@ def iso_cycle' : {f : perm α // is_cycle f} ≃ {s : cycle α // s.nodup ∧ s. notation `c[` l:(foldr `, ` (h t, list.cons h t) list.nil `]`) := cycle.form_perm ↑l (cycle.nodup_coe_iff.mpr dec_trivial) -instance repr_perm [has_repr α] : has_repr (perm α) := +meta instance repr_perm [has_repr α] : has_repr (perm α) := ⟨λ f, repr (multiset.pmap (λ (g : perm α) (hg : g.is_cycle), iso_cycle ⟨g, hg⟩) -- to_cycle is faster? (perm.cycle_factors_finset f).val diff --git a/test/cycle.lean b/test/cycle.lean index 90b5222aa9fb7..bc3424ef5b4bb 100644 --- a/test/cycle.lean +++ b/test/cycle.lean @@ -1,5 +1,5 @@ import data.list.cycle run_cmd guard ("c[1, 4, 3, 2]" = repr (↑[1, 4, 3, 2] : cycle ℕ)) -run_cmd guard ("c[1, 4, 3, 2]" = repr (↑[2, 1, 4, 3] : cycle ℕ)) -run_cmd guard ("c[-1, 2, 1, 4]" = repr (↑[(2 : ℤ), 1, 4, -1] : cycle ℤ)) +run_cmd guard ("c[2, 1, 4, 3]" = repr (↑[2, 1, 4, 3] : cycle ℕ)) +run_cmd guard ("c[2, 1, 4, -1]" = repr (↑[(2 : ℤ), 1, 4, -1] : cycle ℤ)) From 23c61a3ca8d0bc8e9309e0128fb40936d7a2c806 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Mar=C3=ADa=20In=C3=A9s=20de=20Frutos-Fern=C3=A1ndez?= Date: Sat, 14 Jan 2023 06:59:42 +0000 Subject: [PATCH 0240/1291] feat(analysis/normed/group/seminorm): add nonarchimedean (semi)norms (#17851) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We introduce nonarchimedean seminorms and norms on additive groups. Co-authored-by: María Inés de Frutos-Fernández <88536493+mariainesdff@users.noreply.github.com> --- src/analysis/normed/group/seminorm.lean | 279 ++++++++++++++++++++++-- 1 file changed, 266 insertions(+), 13 deletions(-) diff --git a/src/analysis/normed/group/seminorm.lean b/src/analysis/normed/group/seminorm.lean index c84c02e26aee6..5bc0b7f3f7854 100644 --- a/src/analysis/normed/group/seminorm.lean +++ b/src/analysis/normed/group/seminorm.lean @@ -16,9 +16,13 @@ is positive-semidefinite and subadditive. A norm further only maps zero to zero. * `add_group_seminorm`: A function `f` from an additive group `G` to the reals that preserves zero, takes nonnegative values, is subadditive and such that `f (-x) = f x` for all `x`. +* `nonarch_add_group_seminorm`: A function `f` from an additive group `G` to the reals that + preserves zero, takes nonnegative values, is nonarchimedean and such that `f (-x) = f x` + for all `x`. * `group_seminorm`: A function `f` from a group `G` to the reals that sends one to zero, takes nonnegative values, is submultiplicative and such that `f x⁻¹ = f x` for all `x`. * `add_group_norm`: A seminorm `f` such that `f x = 0 → x = 0` for all `x`. +* `nonarch_add_group_norm`: A nonarchimedean seminorm `f` such that `f x = 0 → x = 0` for all `x`. * `group_norm`: A seminorm `f` such that `f x = 0 → x = 1` for all `x`. ## Notes @@ -26,6 +30,10 @@ is positive-semidefinite and subadditive. A norm further only maps zero to zero. The corresponding hom classes are defined in `analysis.order.hom.basic` to be used by absolute values. +We do not define `nonarch_add_group_seminorm` as an extension of `add_group_seminorm` to avoid +having a superfluous `add_le'` field in the resulting structure. The same applies to +`nonarch_add_group_norm`. + ## References * [H. H. Schaefer, *Topological Vector Spaces*][schaefer1966] @@ -44,19 +52,31 @@ variables {ι R R' E F G : Type*} /-- A seminorm on an additive group `G` is a function `f : G → ℝ` that preserves zero, is subadditive and such that `f (-x) = f x` for all `x`. -/ +@[protect_proj] structure add_group_seminorm (G : Type*) [add_group G] extends zero_hom G ℝ := (add_le' : ∀ r s, to_fun (r + s) ≤ to_fun r + to_fun s) (neg' : ∀ r, to_fun (-r) = to_fun r) /-- A seminorm on a group `G` is a function `f : G → ℝ` that sends one to zero, is submultiplicative and such that `f x⁻¹ = f x` for all `x`. -/ -@[to_additive] +@[to_additive, protect_proj] structure group_seminorm (G : Type*) [group G] := (to_fun : G → ℝ) (map_one' : to_fun 1 = 0) (mul_le' : ∀ x y, to_fun (x * y) ≤ to_fun x + to_fun y) (inv' : ∀ x, to_fun x⁻¹ = to_fun x) +/-- A nonarchimedean seminorm on an additive group `G` is a function `f : G → ℝ` that preserves +zero, is nonarchimedean and such that `f (-x) = f x` for all `x`. -/ +@[protect_proj] +structure nonarch_add_group_seminorm (G : Type*) [add_group G] extends zero_hom G ℝ := +(add_le_max' : ∀ r s, to_fun (r + s) ≤ max (to_fun r) (to_fun s)) +(neg' : ∀ r, to_fun (-r) = to_fun r) + +/-! NOTE: We do not define `nonarch_add_group_seminorm` as an extension of `add_group_seminorm` + to avoid having a superfluous `add_le'` field in the resulting structure. The same applies to + `nonarch_add_group_norm` below. -/ + /-- A norm on an additive group `G` is a function `f : G → ℝ` that preserves zero, is subadditive and such that `f (-x) = f x` and `f x = 0 → x = 0` for all `x`. -/ @[protect_proj] @@ -69,11 +89,70 @@ and such that `f x⁻¹ = f x` and `f x = 0 → x = 1` for all `x`. -/ structure group_norm (G : Type*) [group G] extends group_seminorm G := (eq_one_of_map_eq_zero' : ∀ x, to_fun x = 0 → x = 1) -attribute [nolint doc_blame] add_group_seminorm.to_zero_hom add_group_norm.to_add_group_seminorm - group_norm.to_group_seminorm +/-- A nonarchimedean norm on an additive group `G` is a function `f : G → ℝ` that preserves zero, is +nonarchimedean and such that `f (-x) = f x` and `f x = 0 → x = 0` for all `x`. -/ +@[protect_proj] +structure nonarch_add_group_norm (G : Type*) [add_group G] extends nonarch_add_group_seminorm G := +(eq_zero_of_map_eq_zero' : ∀ x, to_fun x = 0 → x = 0) + +attribute [nolint doc_blame] add_group_seminorm.to_zero_hom add_group_norm.to_add_group_seminorm + group_norm.to_group_seminorm nonarch_add_group_seminorm.to_zero_hom + nonarch_add_group_norm.to_nonarch_add_group_seminorm attribute [to_additive] group_norm.to_group_seminorm +/-- `nonarch_add_group_seminorm_class F α` states that `F` is a type of nonarchimedean seminorms on +the additive group `α`. + +You should extend this class when you extend `nonarch_add_group_seminorm`. -/ +@[protect_proj] +class nonarch_add_group_seminorm_class (F : Type*) (α : out_param $ Type*) [add_group α] + extends nonarchimedean_hom_class F α ℝ := +(map_zero (f : F) : f 0 = 0) +(map_neg_eq_map' (f : F) (a : α) : f (-a) = f a) + +/-- `nonarch_add_group_norm_class F α` states that `F` is a type of nonarchimedean norms on the +additive group `α`. + +You should extend this class when you extend `nonarch_add_group_norm`. -/ +@[protect_proj] +class nonarch_add_group_norm_class (F : Type*) (α : out_param $ Type*) [add_group α] + extends nonarch_add_group_seminorm_class F α := +(eq_zero_of_map_eq_zero (f : F) {a : α} : f a = 0 → a = 0) + +section nonarch_add_group_seminorm_class +variables [add_group E] [nonarch_add_group_seminorm_class F E] (f : F) (x y : E) +include E + +lemma map_sub_le_max : f (x - y) ≤ max (f x) (f y) := +by { rw [sub_eq_add_neg, ← nonarch_add_group_seminorm_class.map_neg_eq_map' f y], + exact map_add_le_max _ _ _ } + +end nonarch_add_group_seminorm_class + +@[priority 100] -- See note [lower instance priority] +instance nonarch_add_group_seminorm_class.to_add_group_seminorm_class [add_group E] + [nonarch_add_group_seminorm_class F E] : + add_group_seminorm_class F E ℝ := +{ map_add_le_add := λ f x y, begin + have h_nonneg : ∀ a, 0 ≤ f a, + { intro a, + rw [← nonarch_add_group_seminorm_class.map_zero f, ← sub_self a], + exact le_trans (map_sub_le_max _ _ _) (by rw max_self (f a)) }, + exact le_trans (map_add_le_max _ _ _) + (max_le (le_add_of_nonneg_right (h_nonneg _)) (le_add_of_nonneg_left (h_nonneg _))), + end, + map_neg_eq_map := nonarch_add_group_seminorm_class.map_neg_eq_map', + ..‹nonarch_add_group_seminorm_class F E› } + +@[priority 100] -- See note [lower instance priority] +instance nonarch_add_group_norm_class.to_add_group_norm_class [add_group E] + [nonarch_add_group_norm_class F E] : + add_group_norm_class F E ℝ := +{ map_add_le_add := map_add_le_add, + map_neg_eq_map := nonarch_add_group_seminorm_class.map_neg_eq_map', + ..‹nonarch_add_group_norm_class F E› } + /-! ### Seminorms -/ namespace group_seminorm @@ -90,7 +169,7 @@ variables [group E] [group F] [group G] {p q : group_seminorm E} /-- Helper instance for when there's too many metavariables to apply `fun_like.has_coe_to_fun`. -/ @[to_additive "Helper instance for when there's too many metavariables to apply `fun_like.has_coe_to_fun`. "] -instance : has_coe_to_fun (group_seminorm E) (λ _, E → ℝ) := ⟨to_fun⟩ +instance : has_coe_to_fun (group_seminorm E) (λ _, E → ℝ) := ⟨group_seminorm.to_fun⟩ @[simp, to_additive] lemma to_fun_eq_coe : p.to_fun = p := rfl @@ -102,8 +181,8 @@ partial_order.lift _ fun_like.coe_injective @[to_additive] lemma le_def : p ≤ q ↔ (p : E → ℝ) ≤ q := iff.rfl @[to_additive] lemma lt_def : p < q ↔ (p : E → ℝ) < q := iff.rfl -@[simp, to_additive] lemma coe_le_coe : (p : E → ℝ) ≤ q ↔ p ≤ q := iff.rfl -@[simp, to_additive] lemma coe_lt_coe : (p : E → ℝ) < q ↔ p < q := iff.rfl +@[simp, to_additive, norm_cast] lemma coe_le_coe : (p : E → ℝ) ≤ q ↔ p ≤ q := iff.rfl +@[simp, to_additive, norm_cast] lemma coe_lt_coe : (p : E → ℝ) < q ↔ p < q := iff.rfl variables (p q) (f : F →* E) @@ -113,7 +192,7 @@ variables (p q) (f : F →* E) mul_le' := λ _ _, (zero_add _).ge, inv' := λ x, rfl}⟩ -@[simp, to_additive] lemma coe_zero : ⇑(0 : group_seminorm E) = 0 := rfl +@[simp, to_additive, norm_cast] lemma coe_zero : ⇑(0 : group_seminorm E) = 0 := rfl @[simp, to_additive] lemma zero_apply (x : E) : (0 : group_seminorm E) x = 0 := rfl @[to_additive] instance : inhabited (group_seminorm E) := ⟨0⟩ @@ -141,7 +220,7 @@ variables (p q) (f : F →* E) ((map_mul_le_add q x y).trans $ add_le_add le_sup_right le_sup_right), inv' := λ x, by rw [pi.sup_apply, pi.sup_apply, map_inv_eq_map p, map_inv_eq_map q] }⟩ -@[simp, to_additive] lemma coe_sup : ⇑(p ⊔ q) = p ⊔ q := rfl +@[simp, to_additive, norm_cast] lemma coe_sup : ⇑(p ⊔ q) = p ⊔ q := rfl @[simp, to_additive] lemma sup_apply (x : E) : (p ⊔ q) x = p x ⊔ q x := rfl @[to_additive] instance : semilattice_sup (group_seminorm E) := @@ -245,7 +324,7 @@ instance : has_smul R (add_group_seminorm E) := end, neg' := λ x, by rw map_neg_eq_map }⟩ -@[simp] lemma coe_smul (r : R) (p : add_group_seminorm E) : ⇑(r • p) = r • p := rfl +@[simp, norm_cast] lemma coe_smul (r : R) (p : add_group_seminorm E) : ⇑(r • p) = r • p := rfl @[simp] lemma smul_apply (r : R) (p : add_group_seminorm E) (x : E) : (r • p) x = r • p x := rfl instance [has_smul R' ℝ] [has_smul R' ℝ≥0] [is_scalar_tower R' ℝ≥0 ℝ] @@ -261,6 +340,77 @@ ext $ λ x, real.smul_max _ _ end add_group_seminorm +namespace nonarch_add_group_seminorm +section add_group +variables [add_group E] [add_group F] [add_group G] {p q : nonarch_add_group_seminorm E} + +instance nonarch_add_group_seminorm_class : + nonarch_add_group_seminorm_class (nonarch_add_group_seminorm E) E := +{ coe := λ f, f.to_fun, + coe_injective' := λ f g h, by cases f; cases g; congr', + map_add_le_max := λ f, f.add_le_max', + map_zero := λ f, f.map_zero', + map_neg_eq_map' := λ f, f.neg', } + +/-- Helper instance for when there's too many metavariables to apply `fun_like.has_coe_to_fun`. -/ +instance : has_coe_to_fun (nonarch_add_group_seminorm E) (λ _, E → ℝ) := +⟨nonarch_add_group_seminorm.to_fun⟩ + +@[simp] lemma to_fun_eq_coe : p.to_fun = p := rfl + +@[ext] lemma ext : (∀ x, p x = q x) → p = q := fun_like.ext p q + +noncomputable instance : partial_order (nonarch_add_group_seminorm E) := +partial_order.lift _ fun_like.coe_injective + +lemma le_def : p ≤ q ↔ (p : E → ℝ) ≤ q := iff.rfl +lemma lt_def : p < q ↔ (p : E → ℝ) < q := iff.rfl + +@[simp, norm_cast] lemma coe_le_coe : (p : E → ℝ) ≤ q ↔ p ≤ q := iff.rfl +@[simp, norm_cast] lemma coe_lt_coe : (p : E → ℝ) < q ↔ p < q := iff.rfl + +variables (p q) (f : F →+ E) + +instance : has_zero (nonarch_add_group_seminorm E) := +⟨{ to_fun := 0, + map_zero' := pi.zero_apply _, + add_le_max' := λ r s, by simp only [pi.zero_apply, max_eq_right], + neg' := λ x, rfl}⟩ + +@[simp, norm_cast] lemma coe_zero : ⇑(0 : nonarch_add_group_seminorm E) = 0 := rfl +@[simp] lemma zero_apply (x : E) : (0 : nonarch_add_group_seminorm E) x = 0 := rfl + +instance : inhabited (nonarch_add_group_seminorm E) := ⟨0⟩ + +-- TODO: define `has_Sup` too, from the skeleton at +-- https://github.com/leanprover-community/mathlib/pull/11329#issuecomment-1008915345 +instance : has_sup (nonarch_add_group_seminorm E) := +⟨λ p q, + { to_fun := p ⊔ q, + map_zero' := by rw [pi.sup_apply, ←map_zero p, sup_eq_left, map_zero p, map_zero q], + add_le_max' := λ x y, sup_le + ((map_add_le_max p x y).trans $ max_le_max le_sup_left le_sup_left) + ((map_add_le_max q x y).trans $ max_le_max le_sup_right le_sup_right), + neg' := λ x, by rw [pi.sup_apply, pi.sup_apply, map_neg_eq_map p, map_neg_eq_map q] }⟩ + +@[simp, norm_cast] lemma coe_sup : ⇑(p ⊔ q) = p ⊔ q := rfl +@[simp] lemma sup_apply (x : E) : (p ⊔ q) x = p x ⊔ q x := rfl + +noncomputable instance : semilattice_sup (nonarch_add_group_seminorm E) := +fun_like.coe_injective.semilattice_sup _ coe_sup + +end add_group + +section add_comm_group +variables [add_comm_group E] [add_comm_group F] (p q : nonarch_add_group_seminorm E) (x y : E) + +lemma add_bdd_below_range_add {p q : nonarch_add_group_seminorm E} {x : E} : + bdd_below (range $ λ y, p y + q (x - y)) := +⟨0, by { rintro _ ⟨x, rfl⟩, dsimp, positivity }⟩ + +end add_comm_group +end nonarch_add_group_seminorm + namespace group_seminorm variables [group E] [has_smul R ℝ] [has_smul R ℝ≥0] [is_scalar_tower R ℝ≥0 ℝ] @@ -297,7 +447,7 @@ instance [has_smul R' ℝ] [has_smul R' ℝ≥0] [is_scalar_tower R' ℝ≥0 ℝ [is_scalar_tower R R' ℝ] : is_scalar_tower R R' (group_seminorm E) := ⟨λ r a p, ext $ λ x, smul_assoc r a $ p x⟩ -@[simp, to_additive add_group_seminorm.coe_smul] +@[simp, to_additive add_group_seminorm.coe_smul, norm_cast] lemma coe_smul (r : R) (p : group_seminorm E) : ⇑(r • p) = r • p := rfl @[simp, to_additive add_group_seminorm.smul_apply] @@ -312,6 +462,53 @@ ext $ λ x, real.smul_max _ _ end group_seminorm +namespace nonarch_add_group_seminorm +variables [add_group E] [has_smul R ℝ] [has_smul R ℝ≥0] [is_scalar_tower R ℝ≥0 ℝ] + +instance [decidable_eq E] : has_one (nonarch_add_group_seminorm E) := +⟨{ to_fun := λ x, if x = 0 then 0 else 1, + map_zero' := if_pos rfl, + add_le_max' := λ x y, begin + by_cases hx : x = 0, + { rw [if_pos hx, hx, zero_add], exact le_max_of_le_right (le_refl _) }, + { rw if_neg hx, split_ifs; norm_num } + end, + neg' := λ x, by simp_rw neg_eq_zero }⟩ + +@[simp] lemma apply_one [decidable_eq E] (x : E) : + (1 : nonarch_add_group_seminorm E) x = if x = 0 then 0 else 1 := rfl + +/-- Any action on `ℝ` which factors through `ℝ≥0` applies to a `nonarch_add_group_seminorm`. -/ +instance : has_smul R (nonarch_add_group_seminorm E) := +⟨λ r p, + { to_fun := λ x, r • p x, + map_zero' := by simp only [←smul_one_smul ℝ≥0 r (_ : ℝ), nnreal.smul_def, smul_eq_mul, + map_zero p, mul_zero], + add_le_max' := λ x y, begin + simp only [←smul_one_smul ℝ≥0 r (_ : ℝ), nnreal.smul_def, smul_eq_mul, + ← mul_max_of_nonneg _ _ nnreal.zero_le_coe], + exact mul_le_mul_of_nonneg_left (map_add_le_max p _ _) nnreal.zero_le_coe, + end, + neg' := λ x, by rw map_neg_eq_map p }⟩ + +instance [has_smul R' ℝ] [has_smul R' ℝ≥0] [is_scalar_tower R' ℝ≥0 ℝ] [has_smul R R'] + [is_scalar_tower R R' ℝ] : is_scalar_tower R R' (nonarch_add_group_seminorm E) := +⟨λ r a p, ext $ λ x, smul_assoc r a $ p x⟩ + +@[simp, norm_cast] lemma coe_smul (r : R) (p : nonarch_add_group_seminorm E) : ⇑(r • p) = r • p := +rfl + +@[simp] +lemma smul_apply (r : R) (p : nonarch_add_group_seminorm E) (x : E) : (r • p) x = r • p x := rfl + +lemma smul_sup (r : R) (p q : nonarch_add_group_seminorm E) : r • (p ⊔ q) = r • p ⊔ r • q := +have real.smul_max : ∀ x y : ℝ, r • max x y = max (r • x) (r • y), +from λ x y, by simpa only [←smul_eq_mul, ←nnreal.smul_def, smul_one_smul ℝ≥0 r (_ : ℝ)] + using mul_max_of_nonneg x y (r • 1 : ℝ≥0).prop, +ext $ λ x, real.smul_max _ _ + +end nonarch_add_group_seminorm + /-! ### Norms -/ namespace group_norm @@ -342,8 +539,8 @@ partial_order.lift _ fun_like.coe_injective @[to_additive] lemma le_def : p ≤ q ↔ (p : E → ℝ) ≤ q := iff.rfl @[to_additive] lemma lt_def : p < q ↔ (p : E → ℝ) < q := iff.rfl -@[simp, to_additive] lemma coe_le_coe : (p : E → ℝ) ≤ q ↔ p ≤ q := iff.rfl -@[simp, to_additive] lemma coe_lt_coe : (p : E → ℝ) < q ↔ p < q := iff.rfl +@[simp, to_additive, norm_cast] lemma coe_le_coe : (p : E → ℝ) ≤ q ↔ p ≤ q := iff.rfl +@[simp, to_additive, norm_cast] lemma coe_lt_coe : (p : E → ℝ) < q ↔ p < q := iff.rfl variables (p q) (f : F →* E) @@ -362,7 +559,7 @@ variables (p q) (f : F →* E) lt_sup_iff.2 $ or.inl $ map_pos_of_ne_one p h, ..p.to_group_seminorm ⊔ q.to_group_seminorm }⟩ -@[simp, to_additive] lemma coe_sup : ⇑(p ⊔ q) = p ⊔ q := rfl +@[simp, to_additive, norm_cast] lemma coe_sup : ⇑(p ⊔ q) = p ⊔ q := rfl @[simp, to_additive] lemma sup_apply (x : E) : (p ⊔ q) x = p x ⊔ q x := rfl @[to_additive] instance : semilattice_sup (group_norm E) := @@ -397,3 +594,59 @@ lemma apply_one (x : E) : (1 : group_norm E) x = if x = 1 then 0 else 1 := rfl @[to_additive] instance : inhabited (group_norm E) := ⟨1⟩ end group_norm + +namespace nonarch_add_group_norm +section add_group +variables [add_group E] [add_group F] {p q : nonarch_add_group_norm E} + +instance nonarch_add_group_norm_class : + nonarch_add_group_norm_class (nonarch_add_group_norm E) E := +{ coe := λ f, f.to_fun, + coe_injective' := λ f g h, by cases f; cases g; congr', + map_add_le_max := λ f, f.add_le_max', + map_zero := λ f, f.map_zero', + map_neg_eq_map' := λ f, f.neg', + eq_zero_of_map_eq_zero := λ f, f.eq_zero_of_map_eq_zero' } + +/-- Helper instance for when there's too many metavariables to apply `fun_like.has_coe_to_fun`. -/ +noncomputable instance : has_coe_to_fun (nonarch_add_group_norm E) (λ _, E → ℝ) := +fun_like.has_coe_to_fun + +@[simp] lemma to_fun_eq_coe : p.to_fun = p := rfl + +@[ext] lemma ext : (∀ x, p x = q x) → p = q := fun_like.ext p q + +noncomputable instance : partial_order (nonarch_add_group_norm E) := +partial_order.lift _ fun_like.coe_injective + +lemma le_def : p ≤ q ↔ (p : E → ℝ) ≤ q := iff.rfl +lemma lt_def : p < q ↔ (p : E → ℝ) < q := iff.rfl + +@[simp, norm_cast] lemma coe_le_coe : (p : E → ℝ) ≤ q ↔ p ≤ q := iff.rfl +@[simp, norm_cast] lemma coe_lt_coe : (p : E → ℝ) < q ↔ p < q := iff.rfl + +variables (p q) (f : F →+ E) + +instance : has_sup (nonarch_add_group_norm E) := +⟨λ p q, + { eq_zero_of_map_eq_zero' := λ x hx, of_not_not $ λ h, hx.not_gt $ + lt_sup_iff.2 $ or.inl $ map_pos_of_ne_zero p h, + ..p.to_nonarch_add_group_seminorm ⊔ q.to_nonarch_add_group_seminorm }⟩ + +@[simp, norm_cast] lemma coe_sup : ⇑(p ⊔ q) = p ⊔ q := rfl +@[simp] lemma sup_apply (x : E) : (p ⊔ q) x = p x ⊔ q x := rfl + +noncomputable instance : semilattice_sup (nonarch_add_group_norm E) := +fun_like.coe_injective.semilattice_sup _ coe_sup + +instance [decidable_eq E] : has_one (nonarch_add_group_norm E) := +⟨{ eq_zero_of_map_eq_zero' := λ x, zero_ne_one.ite_eq_left_iff.1, + ..(1 : nonarch_add_group_seminorm E) }⟩ + +@[simp] lemma apply_one [decidable_eq E] (x : E) : + (1 : nonarch_add_group_norm E) x = if x = 0 then 0 else 1 := rfl + +instance [decidable_eq E] : inhabited (nonarch_add_group_norm E) := ⟨1⟩ + +end add_group +end nonarch_add_group_norm From 995b47e555f1b6297c7cf16855f1023e355219fb Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sat, 14 Jan 2023 09:00:00 +0000 Subject: [PATCH 0241/1291] =?UTF-8?q?refactor(*):=20use=20`option.map?= =?UTF-8?q?=E2=82=82`=20(#18081)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Relevant parts are forward-ported as leanprover-community/mathlib4#1439 Co-authored-by: Eric Wieser Co-authored-by: Yaël Dillies --- .../convergents_equiv.lean | 17 ++---- src/algebra/group/with_one/defs.lean | 57 ++++++----------- src/algebra/monoid_algebra/grading.lean | 2 +- src/algebra/order/monoid/with_top.lean | 29 ++------- src/algebra/order/ring/with_top.lean | 31 +++++----- src/data/finset/n_ary.lean | 12 ++++ src/data/option/n_ary.lean | 16 ++++- src/data/seq/seq.lean | 61 +++---------------- src/data/set/n_ary.lean | 12 ++++ src/order/filter/n_ary.lean | 12 ++++ src/order/filter/pointwise.lean | 4 +- src/order/with_bot.lean | 13 ++-- 12 files changed, 114 insertions(+), 152 deletions(-) diff --git a/src/algebra/continued_fractions/convergents_equiv.lean b/src/algebra/continued_fractions/convergents_equiv.lean index d1a1803aceb7d..dee99ac682d10 100644 --- a/src/algebra/continued_fractions/convergents_equiv.lean +++ b/src/algebra/continued_fractions/convergents_equiv.lean @@ -111,7 +111,7 @@ squashed into position `n`. -/ lemma squash_seq_nth_of_not_terminated {gp_n gp_succ_n : pair K} (s_nth_eq : s.nth n = some gp_n) (s_succ_nth_eq : s.nth (n + 1) = some gp_succ_n) : (squash_seq s n).nth n = some ⟨gp_n.a, gp_n.b + gp_succ_n.a / gp_succ_n.b⟩ := -by simp [*, squash_seq, (seq.zip_with_nth_some (seq.nats_nth n) s_nth_eq _)] +by simp [*, squash_seq] /-- The values before the squashed position stay the same. -/ lemma squash_seq_nth_of_lt {m : ℕ} (m_lt_n : m < n) : (squash_seq s n).nth m = s.nth m := @@ -123,8 +123,7 @@ begin s.ge_stable n.le_succ s_succ_nth_eq, obtain ⟨gp_m, s_mth_eq⟩ : ∃ gp_m, s.nth m = some gp_m, from s.ge_stable (le_of_lt m_lt_n) s_nth_eq, - simp [*, squash_seq, (seq.zip_with_nth_some (seq.nats_nth m) s_mth_eq _), - (ne_of_lt m_lt_n)] } + simp [*, squash_seq, m_lt_n.ne] } end /-- Squashing at position `n + 1` and taking the tail is the same as squashing the tail of the @@ -141,19 +140,15 @@ begin { obtain ⟨gp_succ_n, s_succ_nth_eq⟩ : ∃ gp_succ_n, s.nth (n + 1) = some gp_succ_n, from s.ge_stable (n + 1).le_succ s_succ_succ_nth_eq, -- apply extensionality with `m` and continue by cases `m = n`. - ext m, + ext1 m, cases decidable.em (m = n) with m_eq_n m_ne_n, { have : s.tail.nth n = some gp_succ_n, from (s.nth_tail n).trans s_succ_nth_eq, - simp [*, squash_seq, seq.nth_tail, (seq.zip_with_nth_some (seq.nats_nth n) this), - (seq.zip_with_nth_some (seq.nats_nth (n + 1)) s_succ_nth_eq)] }, + simp [*, squash_seq] }, { have : s.tail.nth m = s.nth (m + 1), from s.nth_tail m, cases s_succ_mth_eq : s.nth (m + 1), all_goals { have s_tail_mth_eq, from this.trans s_succ_mth_eq }, - { simp only [*, squash_seq, seq.nth_tail, (seq.zip_with_nth_none' s_succ_mth_eq), - (seq.zip_with_nth_none' s_tail_mth_eq)] }, - { simp [*, squash_seq, seq.nth_tail, - (seq.zip_with_nth_some (seq.nats_nth (m + 1)) s_succ_mth_eq), - (seq.zip_with_nth_some (seq.nats_nth m) s_tail_mth_eq)] } } } + { simp only [*, squash_seq, seq.nth_tail, seq.nth_zip_with, option.map₂_none_right] }, + { simp [*, squash_seq] } } } end /-- The auxiliary function `convergents'_aux` returns the same value for a sequence and the diff --git a/src/algebra/group/with_one/defs.lean b/src/algebra/group/with_one/defs.lean index 813e0abb03ac5..18428357b909c 100644 --- a/src/algebra/group/with_one/defs.lean +++ b/src/algebra/group/with_one/defs.lean @@ -19,6 +19,12 @@ this provides an example of an adjunction is proved in `algebra.category.Mon.adj Another result says that adjoining to a group an element `zero` gives a `group_with_zero`. For more information about these structures (which are not that standard in informal mathematics, see `algebra.group_with_zero.basic`) + +## Implementation notes + +At various points in this file, `id $` is used in at the start of a proof field in a structure. This +ensures that the generated `_proof_1` lemmas are stated in terms of the algebraic operations and +not `option.map`, as the latter does not typecheck once `with_zero`/`with_one` is irreducible. -/ universes u v w @@ -49,7 +55,7 @@ instance [has_mul α] : has_mul (with_one α) := ⟨option.lift_or_get (*)⟩ @[to_additive] instance [has_inv α] : has_inv (with_one α) := ⟨λ a, option.map has_inv.inv a⟩ @[to_additive] instance [has_involutive_inv α] : has_involutive_inv (with_one α) := -{ inv_inv := λ a, (option.map_map _ _ _).trans $ by simp_rw [inv_comp_inv, option.map_id, id], +{ inv_inv := id $ λ a, (option.map_map _ _ _).trans $ by simp_rw [inv_comp_inv, option.map_id, id], ..with_one.has_inv } @[to_additive] instance [has_inv α] : inv_one_class (with_one α) := @@ -112,15 +118,12 @@ protected lemma cases_on {P : with_one α → Prop} : ∀ (x : with_one α), P 1 → (∀ a : α, P a) → P x := option.cases_on --- the `show` statements in the proofs are important, because otherwise the generated lemmas --- `with_one.mul_one_class._proof_{1,2}` have an ill-typed statement after `with_one` is made --- irreducible. @[to_additive] instance [has_mul α] : mul_one_class (with_one α) := { mul := (*), one := (1), - one_mul := show ∀ x : with_one α, 1 * x = x, from (option.lift_or_get_is_left_id _).1, - mul_one := show ∀ x : with_one α, x * 1 = x, from (option.lift_or_get_is_right_id _).1 } + one_mul := id $ (option.lift_or_get_is_left_id _).1, + mul_one := id $ (option.lift_or_get_is_right_id _).1 } @[to_additive] instance [semigroup α] : monoid (with_one α) := @@ -153,49 +156,28 @@ instance [one : has_one α] : has_one (with_zero α) := @[simp, norm_cast] lemma coe_one [has_one α] : ((1 : α) : with_zero α) = 1 := rfl instance [has_mul α] : mul_zero_class (with_zero α) := -{ mul := λ o₁ o₂, o₁.bind (λ a, option.map (λ b, a * b) o₂), - zero_mul := λ a, rfl, - mul_zero := λ a, by cases a; refl, +{ mul := option.map₂ (*), + zero_mul := id $ option.map₂_none_left (*), + mul_zero := id $ option.map₂_none_right (*), ..with_zero.has_zero } @[simp, norm_cast] lemma coe_mul {α : Type u} [has_mul α] {a b : α} : ((a * b : α) : with_zero α) = a * b := rfl -@[simp] lemma zero_mul {α : Type u} [has_mul α] - (a : with_zero α) : 0 * a = 0 := rfl - -@[simp] lemma mul_zero {α : Type u} [has_mul α] - (a : with_zero α) : a * 0 = 0 := by cases a; refl - instance [has_mul α] : no_zero_divisors (with_zero α) := -⟨by { rintro (a|a) (b|b) h, exacts [or.inl rfl, or.inl rfl, or.inr rfl, option.no_confusion h] }⟩ +⟨λ a b, id $ option.map₂_eq_none_iff.1⟩ instance [semigroup α] : semigroup_with_zero (with_zero α) := -{ mul_assoc := λ a b c, match a, b, c with - | none, _, _ := rfl - | some a, none, _ := rfl - | some a, some b, none := rfl - | some a, some b, some c := congr_arg some (mul_assoc _ _ _) - end, +{ mul_assoc := id $ λ _ _ _, option.map₂_assoc mul_assoc, ..with_zero.mul_zero_class } instance [comm_semigroup α] : comm_semigroup (with_zero α) := -{ mul_comm := λ a b, match a, b with - | none, _ := (mul_zero _).symm - | some a, none := rfl - | some a, some b := congr_arg some (mul_comm _ _) - end, +{ mul_comm := id $ λ _ _, option.map₂_comm mul_comm, ..with_zero.semigroup_with_zero } instance [mul_one_class α] : mul_zero_one_class (with_zero α) := -{ one_mul := λ a, match a with - | none := rfl - | some a := congr_arg some $ one_mul _ - end, - mul_one := λ a, match a with - | none := rfl - | some a := congr_arg some $ mul_one _ - end, +{ one_mul := id $ option.map₂_left_identity one_mul, + mul_one := id $ option.map₂_right_identity mul_one, ..with_zero.mul_zero_class, ..with_zero.has_one } @@ -234,7 +216,7 @@ instance [has_inv α] : has_inv (with_zero α) := ⟨λ a, option.map has_inv.in @[simp] lemma inv_zero [has_inv α] : (0 : with_zero α)⁻¹ = 0 := rfl instance [has_involutive_inv α] : has_involutive_inv (with_zero α) := -{ inv_inv := λ a, (option.map_map _ _ _).trans $ by simp_rw [inv_comp_inv, option.map_id, id], +{ inv_inv := id $ λ a, (option.map_map _ _ _).trans $ by simp_rw [inv_comp_inv, option.map_id, id], ..with_zero.has_inv } instance [inv_one_class α] : inv_one_class (with_zero α) := @@ -242,8 +224,7 @@ instance [inv_one_class α] : inv_one_class (with_zero α) := ..with_zero.has_one, ..with_zero.has_inv } -instance [has_div α] : has_div (with_zero α) := -⟨λ o₁ o₂, o₁.bind (λ a, option.map (λ b, a / b) o₂)⟩ +instance [has_div α] : has_div (with_zero α) := ⟨option.map₂ (/)⟩ @[norm_cast] lemma coe_div [has_div α] (a b : α) : ↑(a / b : α) = (a / b : with_zero α) := rfl diff --git a/src/algebra/monoid_algebra/grading.lean b/src/algebra/monoid_algebra/grading.lean index 00881606744f4..aec0305a1a724 100644 --- a/src/algebra/monoid_algebra/grading.lean +++ b/src/algebra/monoid_algebra/grading.lean @@ -180,7 +180,7 @@ graded_algebra.of_alg_hom _ dsimp, rw [decompose_aux_single, direct_sum.coe_alg_hom_of, subtype.coe_mk], end) - (λ i x, by convert (decompose_aux_coe f x : _)) + (λ i x, by rw [decompose_aux_coe f x]) -- Lean can't find this later without us repeating it instance grade_by.decomposition : direct_sum.decomposition (grade_by R f) := diff --git a/src/algebra/order/monoid/with_top.lean b/src/algebra/order/monoid/with_top.lean index c40041e3a9e03..2d35a1f9616ef 100644 --- a/src/algebra/order/monoid/with_top.lean +++ b/src/algebra/order/monoid/with_top.lean @@ -54,7 +54,7 @@ end has_one section has_add variables [has_add α] {a b c d : with_top α} {x y : α} -instance : has_add (with_top α) := ⟨λ o₁ o₂, o₁.bind $ λ a, o₂.map $ (+) a⟩ +instance : has_add (with_top α) := ⟨option.map₂ (+)⟩ @[norm_cast] lemma coe_add : ((x + y : α) : with_top α) = x + y := rfl @[norm_cast] lemma coe_bit0 : ((bit0 x : α) : with_top α) = bit0 x := rfl @@ -201,35 +201,16 @@ end end has_add instance [add_semigroup α] : add_semigroup (with_top α) := -{ add_assoc := begin - repeat { refine with_top.rec_top_coe _ _; try { intro }}; - simp [←with_top.coe_add, add_assoc] - end, +{ add_assoc := λ _ _ _, option.map₂_assoc add_assoc, ..with_top.has_add } instance [add_comm_semigroup α] : add_comm_semigroup (with_top α) := -{ add_comm := - begin - repeat { refine with_top.rec_top_coe _ _; try { intro }}; - simp [←with_top.coe_add, add_comm] - end, +{ add_comm := λ _ _, option.map₂_comm add_comm, ..with_top.add_semigroup } instance [add_zero_class α] : add_zero_class (with_top α) := -{ zero_add := - begin - refine with_top.rec_top_coe _ _, - { simp }, - { intro, - rw [←with_top.coe_zero, ←with_top.coe_add, zero_add] } - end, - add_zero := - begin - refine with_top.rec_top_coe _ _, - { simp }, - { intro, - rw [←with_top.coe_zero, ←with_top.coe_add, add_zero] } - end, +{ zero_add := option.map₂_left_identity zero_add, + add_zero := option.map₂_right_identity add_zero, ..with_top.has_zero, ..with_top.has_add } diff --git a/src/algebra/order/ring/with_top.lean b/src/algebra/order/ring/with_top.lean index dc036fafb540e..b67c04b9a94ce 100644 --- a/src/algebra/order/ring/with_top.lean +++ b/src/algebra/order/ring/with_top.lean @@ -30,12 +30,11 @@ variables [has_zero α] [has_mul α] instance : mul_zero_class (with_top α) := { zero := 0, - mul := λ m n, if m = 0 ∨ n = 0 then 0 else m.bind (λa, n.bind $ λb, ↑(a * b)), + mul := λ m n, if m = 0 ∨ n = 0 then 0 else option.map₂ (*) m n, zero_mul := assume a, if_pos $ or.inl rfl, mul_zero := assume a, if_pos $ or.inr rfl } -lemma mul_def {a b : with_top α} : - a * b = if a = 0 ∨ b = 0 then 0 else a.bind (λa, b.bind $ λb, ↑(a * b)) := rfl +lemma mul_def {a b : with_top α} : a * b = if a = 0 ∨ b = 0 then 0 else option.map₂ (*) a b := rfl @[simp] lemma mul_top {a : with_top α} (h : a ≠ 0) : a * ⊤ = ⊤ := by cases a; simp [mul_def, h]; refl @@ -46,6 +45,14 @@ by cases a; simp [mul_def, h]; refl @[simp] lemma top_mul_top : (⊤ * ⊤ : with_top α) = ⊤ := top_mul top_ne_zero +instance [no_zero_divisors α] : no_zero_divisors (with_top α) := +begin + refine ⟨λ a b h₁, decidable.by_contradiction $ λ h₂, _⟩, + rw [mul_def, if_neg h₂] at h₁, + rcases option.mem_map₂_iff.1 h₁ with ⟨a, b, (rfl : _ = _), (rfl : _ = _), hab⟩, + exact h₂ ((eq_zero_or_eq_zero_of_mul_eq_zero hab).imp (congr_arg some) (congr_arg some)) +end + end has_mul section mul_zero_class @@ -55,7 +62,7 @@ variables [mul_zero_class α] @[norm_cast] lemma coe_mul {a b : α} : (↑(a * b) : with_top α) = a * b := decidable.by_cases (assume : a = 0, by simp [this]) $ assume ha, decidable.by_cases (assume : b = 0, by simp [this]) $ assume hb, -by { simp [*, mul_def], refl } +by { simp [*, mul_def] } lemma mul_coe {b : α} (hb : b ≠ 0) : ∀{a : with_top α}, a * b = a.bind (λa:α, ↑(a * b)) | none := show (if (⊤:with_top α) = 0 ∨ (b:with_top α) = 0 then 0 else ⊤ : with_top α) = ⊤, @@ -114,8 +121,8 @@ instance [mul_zero_one_class α] [nontrivial α] : mul_zero_one_class (with_top begin have : ∀ z, map f z = 0 ↔ z = 0, from λ z, (option.map_injective hf).eq_iff' f.to_zero_hom.with_top_map.map_zero, - rcases eq_or_ne x 0 with rfl|hx, { simp }, - rcases eq_or_ne y 0 with rfl|hy, { simp }, + rcases decidable.eq_or_ne x 0 with rfl|hx, { simp }, + rcases decidable.eq_or_ne y 0 with rfl|hy, { simp }, induction x using with_top.rec_top_coe, { simp [hy, this] }, induction y using with_top.rec_top_coe, { have : (f x : with_top S) ≠ 0, by simpa [hf.eq_iff' (map_zero f)] using hx, @@ -124,10 +131,6 @@ instance [mul_zero_one_class α] [nontrivial α] : mul_zero_one_class (with_top end, .. f.to_zero_hom.with_top_map, .. f.to_monoid_hom.to_one_hom.with_top_map } -instance [mul_zero_class α] [no_zero_divisors α] : no_zero_divisors (with_top α) := -⟨λ a b, by cases a; cases b; dsimp [mul_def]; split_ifs; - simp [*, none_eq_top, some_eq_coe, mul_eq_zero] at *⟩ - instance [semigroup_with_zero α] [no_zero_divisors α] : semigroup_with_zero (with_top α) := { mul := (*), zero := 0, @@ -150,7 +153,7 @@ instance [comm_monoid_with_zero α] [no_zero_divisors α] [nontrivial α] : { mul := (*), zero := 0, mul_comm := λ a b, - by simp only [or_comm, mul_def, option.bind_comm a b, mul_comm], + by simp only [or_comm, mul_def, mul_comm, @option.map₂_comm _ _ _ _ a b _ mul_comm], .. with_top.monoid_with_zero } variables [canonically_ordered_comm_semiring α] @@ -202,7 +205,7 @@ instance : mul_zero_class (with_bot α) := with_top.mul_zero_class lemma mul_def {a b : with_bot α} : - a * b = if a = 0 ∨ b = 0 then 0 else a.bind (λa, b.bind $ λb, ↑(a * b)) := rfl + a * b = if a = 0 ∨ b = 0 then 0 else option.map₂ (*) a b := rfl @[simp] lemma mul_bot {a : with_bot α} (h : a ≠ 0) : a * ⊥ = ⊥ := with_top.mul_top h @@ -220,9 +223,7 @@ section mul_zero_class variables [mul_zero_class α] @[norm_cast] lemma coe_mul {a b : α} : (↑(a * b) : with_bot α) = a * b := -decidable.by_cases (assume : a = 0, by simp [this]) $ assume ha, -decidable.by_cases (assume : b = 0, by simp [this]) $ assume hb, -by { simp [*, mul_def], refl } +with_top.coe_mul lemma mul_coe {b : α} (hb : b ≠ 0) {a : with_bot α} : a * b = a.bind (λa:α, ↑(a * b)) := with_top.mul_coe hb diff --git a/src/data/finset/n_ary.lean b/src/data/finset/n_ary.lean index 2ce67aac5c1c9..7b23ce0dd5208 100644 --- a/src/data/finset/n_ary.lean +++ b/src/data/finset/n_ary.lean @@ -321,4 +321,16 @@ lemma image_image₂_right_anticomm {f : α → β' → γ} {g : β → β'} {f' 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 +/-- If `a` is a left identity for `f : α → β → β`, then `{a}` is a left identity for +`finset.image₂ f`. -/ +lemma image₂_left_identity {f : α → γ → γ} {a : α} (h : ∀ b, f a b = b) (t : finset γ) : + image₂ f {a} t = t := +coe_injective $ by rw [coe_image₂, coe_singleton, set.image2_left_identity h] + +/-- If `b` is a right identity for `f : α → β → α`, then `{b}` is a right identity for +`finset.image₂ f`. -/ +lemma image₂_right_identity {f : γ → β → γ} {b : β} (h : ∀ a, f a b = a) (s : finset γ) : + image₂ f s {b} = s := +by rw [image₂_singleton_right, funext h, image_id'] + end finset diff --git a/src/data/option/n_ary.lean b/src/data/option/n_ary.lean index 007541c0723ca..87dc806f4a2a8 100644 --- a/src/data/option/n_ary.lean +++ b/src/data/option/n_ary.lean @@ -42,8 +42,10 @@ because of the lack of universe polymorphism. -/ lemma map₂_def {α β γ : Type*} (f : α → β → γ) (a : option α) (b : option β) : map₂ f a b = f <$> a <*> b := by cases a; refl -@[simp] lemma map₂_some_some (f : α → β → γ) (a : α) (b : β) : map₂ f (some a) (some b) = f a b := +@[simp] lemma map₂_some_some (f : α → β → γ) (a : α) (b : β) : + map₂ f (some a) (some b) = some (f a b) := rfl + lemma map₂_coe_coe (f : α → β → γ) (a : α) (b : β) : map₂ f a b = f a b := rfl @[simp] lemma map₂_none_left (f : α → β → γ) (b : option β) : map₂ f none b = none := rfl @[simp] lemma map₂_none_right (f : α → β → γ) (a : option α) : map₂ f a none = none := @@ -170,4 +172,16 @@ lemma map_map₂_right_anticomm {f : α → β' → γ} {g : β → β'} {f' : map₂ f a (b.map g) = (map₂ f' b a).map g' := by cases a; cases b; simp [h_right_anticomm] +/-- If `a` is a left identity for a binary operation `f`, then `some a` is a left identity for +`option.map₂ f`. -/ +lemma map₂_left_identity {f : α → β → β} {a : α} (h : ∀ b, f a b = b) (o : option β) : + map₂ f (some a) o = o := +by { cases o, exacts [rfl, congr_arg some (h _)] } + +/-- If `b` is a right identity for a binary operation `f`, then `some b` is a right identity for +`option.map₂ f`. -/ +lemma map₂_right_identity {f : α → β → α} {b : β} (h : ∀ a, f a b = a) (o : option α) : + map₂ f o (some b) = o := +by simp [h, map₂] + end option diff --git a/src/data/seq/seq.lean b/src/data/seq/seq.lean index 003053ca0f475..4571c0c1a4020 100644 --- a/src/data/seq/seq.lean +++ b/src/data/seq/seq.lean @@ -411,62 +411,15 @@ def split_at : ℕ → seq α → list α × seq α section zip_with /-- Combine two sequences with a function -/ -def zip_with (f : α → β → γ) : seq α → seq β → seq γ -| ⟨f₁, a₁⟩ ⟨f₂, a₂⟩ := ⟨λ n, - match f₁ n, f₂ n with - | some a, some b := some (f a b) - | _, _ := none - end, - λ n, begin - induction h1 : f₁ n, - { intro H, simp only [(a₁ h1)], refl }, - induction h2 : f₂ n; dsimp [seq.zip_with._match_1]; intro H, - { rw (a₂ h2), cases f₁ (n + 1); refl }, - { rw [h1, h2] at H, contradiction } - end⟩ +def zip_with (f : α → β → γ) (s₁ : seq α) (s₂ : seq β) : seq γ := +⟨λ n, option.map₂ f (s₁.nth n) (s₂.nth n), λ n hn, + option.map₂_eq_none_iff.2 $ (option.map₂_eq_none_iff.1 hn).imp s₁.2 s₂.2⟩ variables {s : seq α} {s' : seq β} {n : ℕ} -lemma zip_with_nth_some {a : α} {b : β} (s_nth_eq_some : s.nth n = some a) -(s_nth_eq_some' : s'.nth n = some b) (f : α → β → γ) : - (zip_with f s s').nth n = some (f a b) := -begin - cases s with st, - have : st n = some a, from s_nth_eq_some, - cases s' with st', - have : st' n = some b, from s_nth_eq_some', - simp only [zip_with, seq.nth, *] -end - -lemma zip_with_nth_none (s_nth_eq_none : s.nth n = none) (f : α → β → γ) : - (zip_with f s s').nth n = none := -begin - cases s with st, - have : st n = none, from s_nth_eq_none, - cases s' with st', - cases st'_nth_eq : st' n; - simp only [zip_with, seq.nth, *] -end - -lemma zip_with_nth_none' (s'_nth_eq_none : s'.nth n = none) (f : α → β → γ) : - (zip_with f s s').nth n = none := -begin - cases s' with st', - have : st' n = none, from s'_nth_eq_none, - cases s with st, - cases st_nth_eq : st n; - simp only [zip_with, seq.nth, *] -end - -lemma nth_zip_with (f : α → β → γ) (s : seq α) (t : seq β) (n : ℕ) : - nth (zip_with f s t) n = option.bind (nth s n) (λ x, option.map (f x) (nth t n)) := -begin - cases hx : nth s n with x, - { rw [zip_with_nth_none hx, option.none_bind'] }, - cases hy : nth t n with y, - { rw [zip_with_nth_none' hy, option.some_bind', option.map_none'] }, - { rw [zip_with_nth_some hx hy, option.some_bind', option.map_some'] } -end +@[simp] lemma nth_zip_with (f : α → β → γ) (s s' n) : + (zip_with f s s').nth n = option.map₂ f (s.nth n) (s'.nth n) := +rfl end zip_with @@ -474,7 +427,7 @@ end zip_with def zip : seq α → seq β → seq (α × β) := zip_with prod.mk lemma nth_zip (s : seq α) (t : seq β) (n : ℕ) : - nth (zip s t) n = option.bind (nth s n) (λ x, option.map (prod.mk x) (nth t n)) := + nth (zip s t) n = option.map₂ prod.mk (nth s n) (nth t n) := nth_zip_with _ _ _ _ /-- Separate a sequence of pairs into two sequences -/ diff --git a/src/data/set/n_ary.lean b/src/data/set/n_ary.lean index 7ebc1528ff099..47d952d27fd72 100644 --- a/src/data/set/n_ary.lean +++ b/src/data/set/n_ary.lean @@ -313,4 +313,16 @@ lemma image_image2_right_anticomm {f : α → β' → γ} {g : β → β'} {f' : image2 f s (t.image g) = (image2 f' t s).image g' := (image_image2_antidistrib_right $ λ a b, (h_right_anticomm b a).symm).symm +/-- If `a` is a left identity for `f : α → β → β`, then `{a}` is a left identity for +`set.image2 f`. -/ +lemma image2_left_identity {f : α → β → β} {a : α} (h : ∀ b, f a b = b) (t : set β) : + image2 f {a} t = t := +by rw [image2_singleton_left, show f a = id, from funext h, image_id] + +/-- If `b` is a right identity for `f : α → β → α`, then `{b}` is a right identity for +`set.image2 f`. -/ +lemma image2_right_identity {f : α → β → α} {b : β} (h : ∀ a, f a b = a) (s : set α) : + image2 f s {b} = s := +by rw [image2_singleton_right, funext h, image_id'] + end set diff --git a/src/order/filter/n_ary.lean b/src/order/filter/n_ary.lean index 0407044df2f72..70744ed0bacce 100644 --- a/src/order/filter/n_ary.lean +++ b/src/order/filter/n_ary.lean @@ -363,4 +363,16 @@ lemma map_map₂_right_anticomm {m : α → β' → γ} {n : β → β'} {m' : map₂ m f (g.map n) = (map₂ m' g f).map n' := (map_map₂_antidistrib_right $ λ a b, (h_right_anticomm b a).symm).symm +/-- If `a` is a left identity for `f : α → β → β`, then `pure a` is a left identity for +`filter.map₂ f`. -/ +lemma map₂_left_identity {f : α → β → β} {a : α} (h : ∀ b, f a b = b) (l : filter β) : + map₂ f (pure a) l = l := +by rw [map₂_pure_left, show f a = id, from funext h, map_id] + +/-- If `b` is a right identity for `f : α → β → α`, then `pure b` is a right identity for +`filter.map₂ f`. -/ +lemma map₂_right_identity {f : α → β → α} {b : β} (h : ∀ a, f a b = a) (l : filter α) : + map₂ f l (pure b) = l := +by rw [map₂_pure_right, funext h, map_id'] + end filter diff --git a/src/order/filter/pointwise.lean b/src/order/filter/pointwise.lean index a1b5b63a54a6e..4c4dca90e03fc 100644 --- a/src/order/filter/pointwise.lean +++ b/src/order/filter/pointwise.lean @@ -280,8 +280,8 @@ variables [mul_one_class α] [mul_one_class β] protected def 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'] } + one_mul := map₂_left_identity one_mul, + mul_one := map₂_right_identity mul_one } localized "attribute [instance] filter.semigroup filter.add_semigroup filter.comm_semigroup filter.add_comm_semigroup filter.mul_one_class filter.add_zero_class" in pointwise diff --git a/src/order/with_bot.lean b/src/order/with_bot.lean index 93d3cd2093e2c..b0cefb0049796 100644 --- a/src/order/with_bot.lean +++ b/src/order/with_bot.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl -/ import order.bounded_order +import data.option.n_ary /-! # `with_bot`, `with_top` @@ -261,13 +262,13 @@ instance [semilattice_sup α] : semilattice_sup (with_bot α) := lemma coe_sup [semilattice_sup α] (a b : α) : ((a ⊔ b : α) : with_bot α) = a ⊔ b := rfl instance [semilattice_inf α] : semilattice_inf (with_bot α) := -{ inf := λ o₁ o₂, o₁.bind (λ a, o₂.map (λ b, a ⊓ b)), +{ inf := option.map₂ (⊓), inf_le_left := λ o₁ o₂ a ha, begin - simp [map] at ha, rcases ha with ⟨b, rfl, c, rfl, rfl⟩, + rcases option.mem_map₂_iff.1 ha with ⟨a, b, (rfl : _ = _), (rfl : _ = _), rfl⟩, exact ⟨_, rfl, inf_le_left⟩ end, inf_le_right := λ o₁ o₂ a ha, begin - simp [map] at ha, rcases ha with ⟨b, rfl, c, rfl, rfl⟩, + rcases option.mem_map₂_iff.1 ha with ⟨a, b, (rfl : _ = _), (rfl : _ = _), rfl⟩, exact ⟨_, rfl, inf_le_right⟩ end, le_inf := λ o₁ o₂ o₃ h₁ h₂ a ha, begin @@ -746,13 +747,13 @@ instance [semilattice_inf α] : semilattice_inf (with_top α) := lemma coe_inf [semilattice_inf α] (a b : α) : ((a ⊓ b : α) : with_top α) = a ⊓ b := rfl instance [semilattice_sup α] : semilattice_sup (with_top α) := -{ sup := λ o₁ o₂, o₁.bind (λ a, o₂.map (λ b, a ⊔ b)), +{ sup := option.map₂ (⊔), le_sup_left := λ o₁ o₂ a ha, begin - simp [map] at ha, rcases ha with ⟨b, rfl, c, rfl, rfl⟩, + rcases option.mem_map₂_iff.1 ha with ⟨a, b, (rfl : _ = _), (rfl : _ = _), rfl⟩, exact ⟨_, rfl, le_sup_left⟩ end, le_sup_right := λ o₁ o₂ a ha, begin - simp [map] at ha, rcases ha with ⟨b, rfl, c, rfl, rfl⟩, + rcases option.mem_map₂_iff.1 ha with ⟨a, b, (rfl : _ = _), (rfl : _ = _), rfl⟩, exact ⟨_, rfl, le_sup_right⟩ end, sup_le := λ o₁ o₂ o₃ h₁ h₂ a ha, begin From c4926d76bb9c5a4a62ed2f03d998081786132105 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sat, 14 Jan 2023 09:00:02 +0000 Subject: [PATCH 0242/1291] chore(ring_theory/is_tensor_product): golf a slow proof (#18169) This speeds it up by 7x according to `set_option profiler tt`, and is also shorter. Co-authored-by: Junyan Xu --- src/algebra/module/ulift.lean | 1 + src/ring_theory/is_tensor_product.lean | 60 +++++++------------------- src/ring_theory/tensor_product.lean | 3 +- 3 files changed, 19 insertions(+), 45 deletions(-) diff --git a/src/algebra/module/ulift.lean b/src/algebra/module/ulift.lean index 30f010e5c1d8d..1aeb5faba2016 100644 --- a/src/algebra/module/ulift.lean +++ b/src/algebra/module/ulift.lean @@ -128,6 +128,7 @@ instance module' [semiring R] [add_comm_monoid M] [module R M] : module R (ulift /-- The `R`-linear equivalence between `ulift M` and `M`. -/ +@[simps apply symm_apply] def module_equiv [semiring R] [add_comm_monoid M] [module R M] : ulift M ≃ₗ[R] M := { to_fun := ulift.down, inv_fun := ulift.up, diff --git a/src/ring_theory/is_tensor_product.lean b/src/ring_theory/is_tensor_product.lean index 50b6eac2fbaef..dceb63548b86c 100644 --- a/src/ring_theory/is_tensor_product.lean +++ b/src/ring_theory/is_tensor_product.lean @@ -244,53 +244,25 @@ lemma is_base_change.of_lift_unique by exactI ∀ [is_scalar_tower R S Q], by exactI ∀ (g : M →ₗ[R] Q), ∃! (g' : N →ₗ[S] Q), (g'.restrict_scalars R).comp f = g) : is_base_change S f := begin - delta is_base_change is_tensor_product, - obtain ⟨g, hg, hg'⟩ := h (ulift.{v₂} $ S ⊗[R] M) + obtain ⟨g, hg, -⟩ := h (ulift.{v₂} $ S ⊗[R] M) (ulift.module_equiv.symm.to_linear_map.comp $ tensor_product.mk R S M 1), let f' : S ⊗[R] M →ₗ[R] N := _, change function.bijective f', let f'' : S ⊗[R] M →ₗ[S] N, - { refine { map_smul' := λ r x, _, ..f' }, - apply tensor_product.induction_on x, - { simp only [map_zero, smul_zero, linear_map.to_fun_eq_coe] }, - { intros x y, - simp only [algebra.of_id_apply, algebra.id.smul_eq_mul, - alg_hom.to_linear_map_apply, linear_map.mul_apply, tensor_product.lift.tmul', - linear_map.smul_apply, ring_hom.id_apply, module.algebra_map_End_apply, f', - _root_.map_mul, tensor_product.smul_tmul', linear_map.coe_restrict_scalars_eq_coe, - linear_map.flip_apply] }, - { intros x y hx hy, dsimp at hx hy ⊢, simp only [hx, hy, smul_add, map_add] } }, - change function.bijective f'', - split, - { apply function.has_left_inverse.injective, - refine ⟨ulift.module_equiv.to_linear_map.comp g, λ x, _⟩, - apply tensor_product.induction_on x, - { simp only [map_zero] }, - { intros x y, - have := (congr_arg (λ a, x • a) (linear_map.congr_fun hg y)).trans - (ulift.module_equiv.symm.map_smul x _).symm, - apply (ulift.module_equiv : ulift.{v₂} (S ⊗ M) ≃ₗ[S] S ⊗ M) - .to_equiv.apply_eq_iff_eq_symm_apply.mpr, - any_goals { apply_instance }, - simpa only [algebra.of_id_apply, smul_tmul', algebra.id.smul_eq_mul, lift.tmul', - linear_map.coe_restrict_scalars_eq_coe, linear_map.flip_apply, alg_hom.to_linear_map_apply, - module.algebra_map_End_apply, linear_map.smul_apply, linear_map.coe_mk, - linear_map.map_smulₛₗ, mk_apply, mul_one] using this }, - { intros x y hx hy, simp only [map_add, hx, hy] } }, - { apply function.has_right_inverse.surjective, - refine ⟨ulift.module_equiv.to_linear_map.comp g, λ x, _⟩, - obtain ⟨g', hg₁, hg₂⟩ := h (ulift.{max v₁ v₃} N) (ulift.module_equiv.symm.to_linear_map.comp f), - have : g' = ulift.module_equiv.symm.to_linear_map := by { refine (hg₂ _ _).symm, refl }, - subst this, - apply (ulift.module_equiv : ulift.{max v₁ v₃} N ≃ₗ[S] N).symm.injective, - simp_rw [← linear_equiv.coe_to_linear_map, ← linear_map.comp_apply], - congr' 1, - apply hg₂, - ext y, - have := linear_map.congr_fun hg y, - dsimp [ulift.module_equiv] at this ⊢, - rw this, - simp only [lift.tmul, linear_map.coe_restrict_scalars_eq_coe, linear_map.flip_apply, - alg_hom.to_linear_map_apply, _root_.map_one, linear_map.one_apply] } + { refine { to_fun := f', map_smul' := λ s x, + tensor_product.induction_on x _ (λ s' y, smul_assoc s s' _) (λ x y hx hy, _), .. f' }, + { rw [map_zero, smul_zero, map_zero, smul_zero] }, + { rw [smul_add, map_add, map_add, smul_add, hx, hy] } }, + simp_rw [fun_like.ext_iff, linear_map.comp_apply, linear_map.restrict_scalars_apply] at hg, + let fe : S ⊗[R] M ≃ₗ[S] N := + linear_equiv.of_linear f'' (ulift.module_equiv.to_linear_map.comp g) _ _, + { exact fe.bijective }, + { rw ← (linear_map.cancel_left (ulift.module_equiv : ulift.{max v₁ v₃} N ≃ₗ[S] N).symm.injective), + refine (h (ulift.{max v₁ v₃} N) $ ulift.module_equiv.symm.to_linear_map.comp f).unique _ rfl, + { apply_instance }, + ext x, + simp only [linear_map.comp_apply, linear_map.restrict_scalars_apply, hg], + apply one_smul }, + { ext x, change (g $ (1 : S) • f x).down = _, rw [one_smul, hg], refl }, end variable {f} diff --git a/src/ring_theory/tensor_product.lean b/src/ring_theory/tensor_product.lean index 24e33059f459e..203263beb8cd1 100644 --- a/src/ring_theory/tensor_product.lean +++ b/src/ring_theory/tensor_product.lean @@ -74,7 +74,8 @@ lemma smul_eq_lsmul_rtensor (a : A) (x : M ⊗[R] N) : a • x = (lsmul R M a).r Given a linear map `M ⊗[R] N →[A] P`, compose it with the canonical bilinear map `M →[A] N →[R] M ⊗[R] N` to form a bilinear map `M →[A] N →[R] P`. -/ @[simps] def curry (f : (M ⊗[R] N) →ₗ[A] P) : M →ₗ[A] (N →ₗ[R] P) := -{ map_smul' := λ c x, linear_map.ext $ λ y, f.map_smul c (x ⊗ₜ y), +{ to_fun := curry (f.restrict_scalars R), + map_smul' := λ c x, linear_map.ext $ λ y, f.map_smul c (x ⊗ₜ y), .. curry (f.restrict_scalars R) } lemma restrict_scalars_curry (f : (M ⊗[R] N) →ₗ[A] P) : From a0ed43b3a5734e36d5af48cbe3332732988bfdbc Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sat, 14 Jan 2023 12:14:27 +0000 Subject: [PATCH 0243/1291] fix(algebra/order/to_interval_mod): align `to_Ico_div` with regular division (#18105) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The previous definition of `to_Ico_div` and `to_Ico_mod` had `x = to_Ico_mod a hb x - to_Ico_div a hb x • b`; but the usual convention for div and mod is `x = x / b + x % b • b`. This change flips the sign, and adjust all the lemmas accordingly. As a result, we now have the more natural `to_Ico_div a hb x = ⌊(x - a) / b⌋` instead of the previous `to_Ico_div a hb x = -⌊(x - a) / b⌋`. This also changes ```diff def mem_Ioo_mod (b x : α) : Prop := -∃ z : ℤ, x + z • b ∈ set.Ioo a (a + b) +∃ z : ℤ, x - z • b ∈ set.Ioo a (a + b) ``` even though the two are equivalent; because this means that `z` corresponds to `to_Ico_div `. --- src/algebra/order/to_interval_mod.lean | 315 +++++++++--------- .../special_functions/complex/arg.lean | 4 +- .../trigonometric/angle.lean | 4 +- src/topology/instances/add_circle.lean | 8 +- 4 files changed, 164 insertions(+), 167 deletions(-) diff --git a/src/algebra/order/to_interval_mod.lean b/src/algebra/order/to_interval_mod.lean index 3a7b1cb9762a4..88b61cd9a7006 100644 --- a/src/algebra/order/to_interval_mod.lean +++ b/src/algebra/order/to_interval_mod.lean @@ -18,10 +18,10 @@ interval. ## Main definitions * `to_Ico_div a hb x` (where `hb : 0 < b`): The unique integer such that this multiple of `b`, - added to `x`, is in `Ico a (a + b)`. + subtracted from `x`, is in `Ico a (a + b)`. * `to_Ico_mod a hb x` (where `hb : 0 < b`): Reduce `x` to the interval `Ico a (a + b)`. * `to_Ioc_div a hb x` (where `hb : 0 < b`): The unique integer such that this multiple of `b`, - added to `x`, is in `Ioc a (a + b)`. + subtracted from `x`, is in `Ioc a (a + b)`. * `to_Ioc_mod a hb x` (where `hb : 0 < b`): Reduce `x` to the interval `Ioc a (a + b)`. -/ @@ -33,39 +33,41 @@ section linear_ordered_add_comm_group variables {α : Type*} [linear_ordered_add_comm_group α] [hα : archimedean α] include hα -/-- The unique integer such that this multiple of `b`, added to `x`, is in `Ico a (a + b)`. -/ +/-- +The unique integer such that this multiple of `b`, subtracted from `x`, is in `Ico a (a + b)`. -/ def to_Ico_div (a : α) {b : α} (hb : 0 < b) (x : α) : ℤ := -(exists_unique_add_zsmul_mem_Ico hb x a).some +(exists_unique_sub_zsmul_mem_Ico hb x a).some -lemma add_to_Ico_div_zsmul_mem_Ico (a : α) {b : α} (hb : 0 < b) (x : α) : - x + to_Ico_div a hb x • b ∈ set.Ico a (a + b) := -(exists_unique_add_zsmul_mem_Ico hb x a).some_spec.1 +lemma sub_to_Ico_div_zsmul_mem_Ico (a : α) {b : α} (hb : 0 < b) (x : α) : + x - to_Ico_div a hb x • b ∈ set.Ico a (a + b) := +(exists_unique_sub_zsmul_mem_Ico hb x a).some_spec.1 -lemma eq_to_Ico_div_of_add_zsmul_mem_Ico {a b x : α} (hb : 0 < b) {y : ℤ} - (hy : x + y • b ∈ set.Ico a (a + b)) : y = to_Ico_div a hb x := -(exists_unique_add_zsmul_mem_Ico hb x a).some_spec.2 y hy +lemma eq_to_Ico_div_of_sub_zsmul_mem_Ico {a b x : α} (hb : 0 < b) {y : ℤ} + (hy : x - y • b ∈ set.Ico a (a + b)) : y = to_Ico_div a hb x := +(exists_unique_sub_zsmul_mem_Ico hb x a).some_spec.2 y hy -/-- The unique integer such that this multiple of `b`, added to `x`, is in `Ioc a (a + b)`. -/ +/-- +The unique integer such that this multiple of `b`, subtracted from `x`, is in `Ioc a (a + b)`. -/ def to_Ioc_div (a : α) {b : α} (hb : 0 < b) (x : α) : ℤ := -(exists_unique_add_zsmul_mem_Ioc hb x a).some +(exists_unique_sub_zsmul_mem_Ioc hb x a).some -lemma add_to_Ioc_div_zsmul_mem_Ioc (a : α) {b : α} (hb : 0 < b) (x : α) : - x + to_Ioc_div a hb x • b ∈ set.Ioc a (a + b) := -(exists_unique_add_zsmul_mem_Ioc hb x a).some_spec.1 +lemma sub_to_Ioc_div_zsmul_mem_Ioc (a : α) {b : α} (hb : 0 < b) (x : α) : + x - to_Ioc_div a hb x • b ∈ set.Ioc a (a + b) := +(exists_unique_sub_zsmul_mem_Ioc hb x a).some_spec.1 -lemma eq_to_Ioc_div_of_add_zsmul_mem_Ioc {a b x : α} (hb : 0 < b) {y : ℤ} - (hy : x + y • b ∈ set.Ioc a (a + b)) : y = to_Ioc_div a hb x := -(exists_unique_add_zsmul_mem_Ioc hb x a).some_spec.2 y hy +lemma eq_to_Ioc_div_of_sub_zsmul_mem_Ioc {a b x : α} (hb : 0 < b) {y : ℤ} + (hy : x - y • b ∈ set.Ioc a (a + b)) : y = to_Ioc_div a hb x := +(exists_unique_sub_zsmul_mem_Ioc hb x a).some_spec.2 y hy /-- Reduce `x` to the interval `Ico a (a + b)`. -/ -def to_Ico_mod (a : α) {b : α} (hb : 0 < b) (x : α) : α := x + to_Ico_div a hb x • b +def to_Ico_mod (a : α) {b : α} (hb : 0 < b) (x : α) : α := x - to_Ico_div a hb x • b /-- Reduce `x` to the interval `Ioc a (a + b)`. -/ -def to_Ioc_mod (a : α) {b : α} (hb : 0 < b) (x : α) : α := x + to_Ioc_div a hb x • b +def to_Ioc_mod (a : α) {b : α} (hb : 0 < b) (x : α) : α := x - to_Ioc_div a hb x • b lemma to_Ico_mod_mem_Ico (a : α) {b : α} (hb : 0 < b) (x : α) : to_Ico_mod a hb x ∈ set.Ico a (a + b) := -add_to_Ico_div_zsmul_mem_Ico a hb x +sub_to_Ico_div_zsmul_mem_Ico a hb x lemma to_Ico_mod_mem_Ico' {b : α} (hb : 0 < b) (x : α) : to_Ico_mod 0 hb x ∈ set.Ico 0 b := @@ -73,7 +75,7 @@ by { convert to_Ico_mod_mem_Ico 0 hb x, exact (zero_add b).symm, } lemma to_Ioc_mod_mem_Ioc (a : α) {b : α} (hb : 0 < b) (x : α) : to_Ioc_mod a hb x ∈ set.Ioc a (a + b) := -add_to_Ioc_div_zsmul_mem_Ioc a hb x +sub_to_Ioc_div_zsmul_mem_Ioc a hb x lemma left_le_to_Ico_mod (a : α) {b : α} (hb : 0 < b) (x : α) : a ≤ to_Ico_mod a hb x := (set.mem_Ico.1 (to_Ico_mod_mem_Ico a hb x)).1 @@ -87,92 +89,92 @@ lemma to_Ico_mod_lt_right (a : α) {b : α} (hb : 0 < b) (x : α) : to_Ico_mod a lemma to_Ioc_mod_le_right (a : α) {b : α} (hb : 0 < b) (x : α) : to_Ioc_mod a hb x ≤ a + b := (set.mem_Ioc.1 (to_Ioc_mod_mem_Ioc a hb x)).2 -@[simp] lemma self_add_to_Ico_div_zsmul (a : α) {b : α} (hb : 0 < b) (x : α) : - x + to_Ico_div a hb x • b = to_Ico_mod a hb x := +@[simp] lemma self_sub_to_Ico_div_zsmul (a : α) {b : α} (hb : 0 < b) (x : α) : + x - to_Ico_div a hb x • b = to_Ico_mod a hb x := rfl -@[simp] lemma self_add_to_Ioc_div_zsmul (a : α) {b : α} (hb : 0 < b) (x : α) : - x + to_Ioc_div a hb x • b = to_Ioc_mod a hb x := +@[simp] lemma self_sub_to_Ioc_div_zsmul (a : α) {b : α} (hb : 0 < b) (x : α) : + x - to_Ioc_div a hb x • b = to_Ioc_mod a hb x := rfl -@[simp] lemma to_Ico_div_zsmul_add_self (a : α) {b : α} (hb : 0 < b) (x : α) : - to_Ico_div a hb x • b + x = to_Ico_mod a hb x := -by rw [add_comm, to_Ico_mod] +@[simp] lemma to_Ico_div_zsmul_sub_self (a : α) {b : α} (hb : 0 < b) (x : α) : + to_Ico_div a hb x • b - x = -to_Ico_mod a hb x := +by rw [to_Ico_mod, neg_sub] -@[simp] lemma to_Ioc_div_zsmul_add_self (a : α) {b : α} (hb : 0 < b) (x : α) : - to_Ioc_div a hb x • b + x = to_Ioc_mod a hb x := -by rw [add_comm, to_Ioc_mod] +@[simp] lemma to_Ioc_div_zsmul_sub_self (a : α) {b : α} (hb : 0 < b) (x : α) : + to_Ioc_div a hb x • b - x = -to_Ioc_mod a hb x := +by rw [to_Ioc_mod, neg_sub] @[simp] lemma to_Ico_mod_sub_self (a : α) {b : α} (hb : 0 < b) (x : α) : - to_Ico_mod a hb x - x = to_Ico_div a hb x • b := -by rw [to_Ico_mod, add_sub_cancel'] + to_Ico_mod a hb x - x = -to_Ico_div a hb x • b := +by rw [to_Ico_mod, sub_sub_cancel_left, neg_smul] @[simp] lemma to_Ioc_mod_sub_self (a : α) {b : α} (hb : 0 < b) (x : α) : - to_Ioc_mod a hb x - x = to_Ioc_div a hb x • b := -by rw [to_Ioc_mod, add_sub_cancel'] + to_Ioc_mod a hb x - x = -to_Ioc_div a hb x • b := +by rw [to_Ioc_mod, sub_sub_cancel_left, neg_smul] @[simp] lemma self_sub_to_Ico_mod (a : α) {b : α} (hb : 0 < b) (x : α) : - x - to_Ico_mod a hb x = -to_Ico_div a hb x • b := -by rw [to_Ico_mod, sub_add_cancel', neg_smul] + x - to_Ico_mod a hb x = to_Ico_div a hb x • b := +by rw [to_Ico_mod, sub_sub_cancel] @[simp] lemma self_sub_to_Ioc_mod (a : α) {b : α} (hb : 0 < b) (x : α) : - x - to_Ioc_mod a hb x = -to_Ioc_div a hb x • b := -by rw [to_Ioc_mod, sub_add_cancel', neg_smul] + x - to_Ioc_mod a hb x = to_Ioc_div a hb x • b := +by rw [to_Ioc_mod, sub_sub_cancel] -@[simp] lemma to_Ico_mod_sub_to_Ico_div_zsmul (a : α) {b : α} (hb : 0 < b) (x : α) : - to_Ico_mod a hb x - to_Ico_div a hb x • b = x := -by rw [to_Ico_mod, add_sub_cancel] +@[simp] lemma to_Ico_mod_add_to_Ico_div_zsmul (a : α) {b : α} (hb : 0 < b) (x : α) : + to_Ico_mod a hb x + to_Ico_div a hb x • b = x := +by rw [to_Ico_mod, sub_add_cancel] -@[simp] lemma to_Ioc_mod_sub_to_Ioc_div_zsmul (a : α) {b : α} (hb : 0 < b) (x : α) : - to_Ioc_mod a hb x - to_Ioc_div a hb x • b = x := -by rw [to_Ioc_mod, add_sub_cancel] +@[simp] lemma to_Ioc_mod_add_to_Ioc_div_zsmul (a : α) {b : α} (hb : 0 < b) (x : α) : + to_Ioc_mod a hb x + to_Ioc_div a hb x • b = x := +by rw [to_Ioc_mod, sub_add_cancel] @[simp] lemma to_Ico_div_zsmul_sub_to_Ico_mod (a : α) {b : α} (hb : 0 < b) (x : α) : - to_Ico_div a hb x • b - to_Ico_mod a hb x = -x := -by rw [←neg_sub, to_Ico_mod_sub_to_Ico_div_zsmul] + to_Ico_div a hb x • b + to_Ico_mod a hb x = x := +by rw [add_comm, to_Ico_mod_add_to_Ico_div_zsmul] @[simp] lemma to_Ioc_div_zsmul_sub_to_Ioc_mod (a : α) {b : α} (hb : 0 < b) (x : α) : - to_Ioc_div a hb x • b - to_Ioc_mod a hb x = -x := -by rw [←neg_sub, to_Ioc_mod_sub_to_Ioc_div_zsmul] + to_Ioc_div a hb x • b + to_Ioc_mod a hb x = x := +by rw [add_comm, to_Ioc_mod_add_to_Ioc_div_zsmul] lemma to_Ico_mod_eq_iff {a b x y : α} (hb : 0 < b) : - to_Ico_mod a hb x = y ↔ y ∈ set.Ico a (a + b) ∧ ∃ z : ℤ, y - x = z • b := + to_Ico_mod a hb x = y ↔ y ∈ set.Ico a (a + b) ∧ ∃ z : ℤ, x = y + z • b := begin refine ⟨λ h, ⟨h ▸ to_Ico_mod_mem_Ico a hb x, to_Ico_div a hb x, - h ▸ to_Ico_mod_sub_self a hb x⟩, + h ▸ (to_Ico_mod_add_to_Ico_div_zsmul _ _ _).symm⟩, λ h, _⟩, rcases h with ⟨hy, z, hz⟩, - rw sub_eq_iff_eq_add' at hz, + rw ←sub_eq_iff_eq_add at hz, subst hz, - rw eq_to_Ico_div_of_add_zsmul_mem_Ico hb hy, + rw eq_to_Ico_div_of_sub_zsmul_mem_Ico hb hy, refl end lemma to_Ioc_mod_eq_iff {a b x y : α} (hb : 0 < b) : - to_Ioc_mod a hb x = y ↔ y ∈ set.Ioc a (a + b) ∧ ∃ z : ℤ, y - x = z • b := + to_Ioc_mod a hb x = y ↔ y ∈ set.Ioc a (a + b) ∧ ∃ z : ℤ, x = y + z • b := begin refine ⟨λ h, ⟨h ▸ to_Ioc_mod_mem_Ioc a hb x, to_Ioc_div a hb x, - h ▸ to_Ioc_mod_sub_self a hb x⟩, + h ▸ (to_Ioc_mod_add_to_Ioc_div_zsmul _ hb _).symm⟩, λ h, _⟩, rcases h with ⟨hy, z, hz⟩, - rw sub_eq_iff_eq_add' at hz, + rw ←sub_eq_iff_eq_add at hz, subst hz, - rw eq_to_Ioc_div_of_add_zsmul_mem_Ioc hb hy, + rw eq_to_Ioc_div_of_sub_zsmul_mem_Ioc hb hy, refl end @[simp] lemma to_Ico_div_apply_left (a : α) {b : α} (hb : 0 < b) : to_Ico_div a hb a = 0 := begin - refine (eq_to_Ico_div_of_add_zsmul_mem_Ico hb _).symm, + refine (eq_to_Ico_div_of_sub_zsmul_mem_Ico hb _).symm, simp [hb] end -@[simp] lemma to_Ioc_div_apply_left (a : α) {b : α} (hb : 0 < b) : to_Ioc_div a hb a = 1 := +@[simp] lemma to_Ioc_div_apply_left (a : α) {b : α} (hb : 0 < b) : to_Ioc_div a hb a = -1 := begin - refine (eq_to_Ioc_div_of_add_zsmul_mem_Ioc hb _).symm, - simp [hb] + refine (eq_to_Ioc_div_of_sub_zsmul_mem_Ioc hb _).symm, + simp [hb], end @[simp] lemma to_Ico_mod_apply_left (a : α) {b : α} (hb : 0 < b) : to_Ico_mod a hb a = a := @@ -185,28 +187,28 @@ end @[simp] lemma to_Ioc_mod_apply_left (a : α) {b : α} (hb : 0 < b) : to_Ioc_mod a hb a = a + b := begin rw [to_Ioc_mod_eq_iff hb, set.right_mem_Ioc], - refine ⟨lt_add_of_pos_right _ hb, 1, _⟩, + refine ⟨lt_add_of_pos_right _ hb, -1, _⟩, simp end lemma to_Ico_div_apply_right (a : α) {b : α} (hb : 0 < b) : - to_Ico_div a hb (a + b) = -1 := + to_Ico_div a hb (a + b) = 1 := begin - refine (eq_to_Ico_div_of_add_zsmul_mem_Ico hb _).symm, + refine (eq_to_Ico_div_of_sub_zsmul_mem_Ico hb _).symm, simp [hb] end lemma to_Ioc_div_apply_right (a : α) {b : α} (hb : 0 < b) : to_Ioc_div a hb (a + b) = 0 := begin - refine (eq_to_Ioc_div_of_add_zsmul_mem_Ioc hb _).symm, + refine (eq_to_Ioc_div_of_sub_zsmul_mem_Ioc hb _).symm, simp [hb] end lemma to_Ico_mod_apply_right (a : α) {b : α} (hb : 0 < b) : to_Ico_mod a hb (a + b) = a := begin rw [to_Ico_mod_eq_iff hb, set.left_mem_Ico], - refine ⟨lt_add_of_pos_right _ hb, -1, _⟩, + refine ⟨lt_add_of_pos_right _ hb, 1, _⟩, simp end @@ -219,68 +221,68 @@ begin end @[simp] lemma to_Ico_div_add_zsmul (a : α) {b : α} (hb : 0 < b) (x : α) (m : ℤ) : - to_Ico_div a hb (x + m • b) = to_Ico_div a hb x - m := + to_Ico_div a hb (x + m • b) = to_Ico_div a hb x + m := begin - refine (eq_to_Ico_div_of_add_zsmul_mem_Ico hb _).symm, - convert add_to_Ico_div_zsmul_mem_Ico a hb x using 1, - simp [sub_smul] + refine (eq_to_Ico_div_of_sub_zsmul_mem_Ico hb _).symm, + convert sub_to_Ico_div_zsmul_mem_Ico a hb x using 1, + simp [add_smul], end @[simp] lemma to_Ioc_div_add_zsmul (a : α) {b : α} (hb : 0 < b) (x : α) (m : ℤ) : - to_Ioc_div a hb (x + m • b) = to_Ioc_div a hb x - m := + to_Ioc_div a hb (x + m • b) = to_Ioc_div a hb x + m := begin - refine (eq_to_Ioc_div_of_add_zsmul_mem_Ioc hb _).symm, - convert add_to_Ioc_div_zsmul_mem_Ioc a hb x using 1, - simp [sub_smul] + refine (eq_to_Ioc_div_of_sub_zsmul_mem_Ioc hb _).symm, + convert sub_to_Ioc_div_zsmul_mem_Ioc a hb x using 1, + simp [add_smul] end @[simp] lemma to_Ico_div_zsmul_add (a : α) {b : α} (hb : 0 < b) (x : α) (m : ℤ) : - to_Ico_div a hb (m • b + x) = to_Ico_div a hb x - m := -by rw [add_comm, to_Ico_div_add_zsmul] + to_Ico_div a hb (m • b + x) = m + to_Ico_div a hb x := +by rw [add_comm, to_Ico_div_add_zsmul, add_comm] @[simp] lemma to_Ioc_div_zsmul_add (a : α) {b : α} (hb : 0 < b) (x : α) (m : ℤ) : - to_Ioc_div a hb (m • b + x) = to_Ioc_div a hb x - m := -by rw [add_comm, to_Ioc_div_add_zsmul] + to_Ioc_div a hb (m • b + x) = to_Ioc_div a hb x + m := +by rw [add_comm, to_Ioc_div_add_zsmul, add_comm] @[simp] lemma to_Ico_div_sub_zsmul (a : α) {b : α} (hb : 0 < b) (x : α) (m : ℤ) : - to_Ico_div a hb (x - m • b) = to_Ico_div a hb x + m := -by rw [sub_eq_add_neg, ←neg_smul, to_Ico_div_add_zsmul, sub_neg_eq_add] + to_Ico_div a hb (x - m • b) = to_Ico_div a hb x - m := +by rw [sub_eq_add_neg, ←neg_smul, to_Ico_div_add_zsmul, sub_eq_add_neg] @[simp] lemma to_Ioc_div_sub_zsmul (a : α) {b : α} (hb : 0 < b) (x : α) (m : ℤ) : - to_Ioc_div a hb (x - m • b) = to_Ioc_div a hb x + m := -by rw [sub_eq_add_neg, ←neg_smul, to_Ioc_div_add_zsmul, sub_neg_eq_add] + to_Ioc_div a hb (x - m • b) = to_Ioc_div a hb x - m := +by rw [sub_eq_add_neg, ←neg_smul, to_Ioc_div_add_zsmul, sub_eq_add_neg] @[simp] lemma to_Ico_div_add_right (a : α) {b : α} (hb : 0 < b) (x : α) : - to_Ico_div a hb (x + b) = to_Ico_div a hb x - 1 := + to_Ico_div a hb (x + b) = to_Ico_div a hb x + 1 := begin convert to_Ico_div_add_zsmul a hb x 1, exact (one_zsmul _).symm end @[simp] lemma to_Ioc_div_add_right (a : α) {b : α} (hb : 0 < b) (x : α) : - to_Ioc_div a hb (x + b) = to_Ioc_div a hb x - 1 := + to_Ioc_div a hb (x + b) = to_Ioc_div a hb x + 1 := begin convert to_Ioc_div_add_zsmul a hb x 1, exact (one_zsmul _).symm end @[simp] lemma to_Ico_div_add_left (a : α) {b : α} (hb : 0 < b) (x : α) : - to_Ico_div a hb (b + x) = to_Ico_div a hb x - 1 := + to_Ico_div a hb (b + x) = to_Ico_div a hb x + 1 := by rw [add_comm, to_Ico_div_add_right] @[simp] lemma to_Ioc_div_add_left (a : α) {b : α} (hb : 0 < b) (x : α) : - to_Ioc_div a hb (b + x) = to_Ioc_div a hb x - 1 := + to_Ioc_div a hb (b + x) = to_Ioc_div a hb x + 1 := by rw [add_comm, to_Ioc_div_add_right] @[simp] lemma to_Ico_div_sub (a : α) {b : α} (hb : 0 < b) (x : α) : - to_Ico_div a hb (x - b) = to_Ico_div a hb x + 1 := + to_Ico_div a hb (x - b) = to_Ico_div a hb x - 1 := begin convert to_Ico_div_sub_zsmul a hb x 1, exact (one_zsmul _).symm end @[simp] lemma to_Ioc_div_sub (a : α) {b : α} (hb : 0 < b) (x : α) : - to_Ioc_div a hb (x - b) = to_Ioc_div a hb x + 1 := + to_Ioc_div a hb (x - b) = to_Ioc_div a hb x - 1 := begin convert to_Ioc_div_sub_zsmul a hb x 1, exact (one_zsmul _).symm @@ -290,22 +292,18 @@ lemma to_Ico_div_sub' (a : α) {b : α} (hb : 0 < b) (x y : α) : to_Ico_div a hb (x - y) = to_Ico_div (a + y) hb x := begin rw eq_comm, - apply eq_to_Ico_div_of_add_zsmul_mem_Ico, - rw sub_add_eq_add_sub, - obtain ⟨hc, ho⟩ := add_to_Ico_div_zsmul_mem_Ico (a + y) hb x, - rw add_right_comm at ho, - exact ⟨le_sub_iff_add_le.mpr hc, sub_lt_iff_lt_add.mpr ho⟩, + apply eq_to_Ico_div_of_sub_zsmul_mem_Ico, + rw [←sub_right_comm, set.sub_mem_Ico_iff_left, add_right_comm], + exact sub_to_Ico_div_zsmul_mem_Ico (a + y) hb x, end lemma to_Ioc_div_sub' (a : α) {b : α} (hb : 0 < b) (x y : α) : to_Ioc_div a hb (x - y) = to_Ioc_div (a + y) hb x := begin rw eq_comm, - apply eq_to_Ioc_div_of_add_zsmul_mem_Ioc, - rw sub_add_eq_add_sub, - obtain ⟨ho, hc⟩ := add_to_Ioc_div_zsmul_mem_Ioc (a + y) hb x, - rw add_right_comm at hc, - exact ⟨lt_sub_iff_add_lt.mpr ho, sub_le_iff_le_add.mpr hc⟩, + apply eq_to_Ioc_div_of_sub_zsmul_mem_Ioc, + rw [←sub_right_comm, set.sub_mem_Ioc_iff_left, add_right_comm], + exact sub_to_Ioc_div_zsmul_mem_Ioc (a + y) hb x, end lemma to_Ico_div_add_right' (a : α) {b : α} (hb : 0 < b) (x y : α) : @@ -317,34 +315,34 @@ lemma to_Ioc_div_add_right' (a : α) {b : α} (hb : 0 < b) (x y : α) : by rw [←sub_neg_eq_add, to_Ioc_div_sub', sub_eq_add_neg] lemma to_Ico_div_neg (a : α) {b : α} (hb : 0 < b) (x : α) : - to_Ico_div a hb (-x) = 1 - to_Ioc_div (-a) hb x := + to_Ico_div a hb (-x) = -(to_Ioc_div (-a) hb x + 1) := begin suffices : to_Ico_div a hb (-x) = -(to_Ioc_div (-(a + b)) hb x), - { rwa [neg_add, ←sub_eq_add_neg, ←to_Ioc_div_add_right', to_Ioc_div_add_right, neg_sub] at this }, + { rwa [neg_add, ←sub_eq_add_neg, ←to_Ioc_div_add_right', to_Ioc_div_add_right] at this }, rw [eq_neg_iff_eq_neg, eq_comm], - apply eq_to_Ioc_div_of_add_zsmul_mem_Ioc, - obtain ⟨hc, ho⟩ := add_to_Ico_div_zsmul_mem_Ico a hb (-x), - rw [←neg_lt_neg_iff, neg_add (-x), neg_neg, ←neg_smul] at ho, - rw [←neg_le_neg_iff, neg_add (-x), neg_neg, ←neg_smul] at hc, + apply eq_to_Ioc_div_of_sub_zsmul_mem_Ioc, + obtain ⟨hc, ho⟩ := sub_to_Ico_div_zsmul_mem_Ico a hb (-x), + rw [←neg_lt_neg_iff, neg_sub' (-x), neg_neg, ←neg_smul] at ho, + rw [←neg_le_neg_iff, neg_sub' (-x), neg_neg, ←neg_smul] at hc, refine ⟨ho, hc.trans_eq _⟩, rw [neg_add, neg_add_cancel_right], end lemma to_Ioc_div_neg (a : α) {b : α} (hb : 0 < b) (x : α) : - to_Ioc_div a hb (-x) = 1 - to_Ico_div (-a) hb x := -by rw [←neg_neg x, to_Ico_div_neg, neg_neg, neg_neg, sub_sub_cancel] + to_Ioc_div a hb (-x) = -(to_Ico_div (-a) hb x + 1) := +by rw [←neg_neg x, to_Ico_div_neg, neg_neg, neg_neg, neg_add', neg_neg, add_sub_cancel] @[simp] lemma to_Ico_mod_add_zsmul (a : α) {b : α} (hb : 0 < b) (x : α) (m : ℤ) : to_Ico_mod a hb (x + m • b) = to_Ico_mod a hb x := begin - rw [to_Ico_mod, to_Ico_div_add_zsmul, to_Ico_mod, sub_smul], + rw [to_Ico_mod, to_Ico_div_add_zsmul, to_Ico_mod, add_smul], abel end @[simp] lemma to_Ioc_mod_add_zsmul (a : α) {b : α} (hb : 0 < b) (x : α) (m : ℤ) : to_Ioc_mod a hb (x + m • b) = to_Ioc_mod a hb x := begin - rw [to_Ioc_mod, to_Ioc_div_add_zsmul, to_Ioc_mod, sub_smul], + rw [to_Ioc_mod, to_Ioc_div_add_zsmul, to_Ioc_mod, add_smul], abel end @@ -402,40 +400,40 @@ end lemma to_Ico_mod_sub' (a : α) {b : α} (hb : 0 < b) (x y : α) : to_Ico_mod a hb (x - y) = to_Ico_mod (a + y) hb x - y := -by simp_rw [to_Ico_mod, to_Ico_div_sub', sub_add_eq_add_sub] +by simp_rw [to_Ico_mod, to_Ico_div_sub', sub_right_comm] lemma to_Ioc_mod_sub' (a : α) {b : α} (hb : 0 < b) (x y : α) : to_Ioc_mod a hb (x - y) = to_Ioc_mod (a + y) hb x - y := -by simp_rw [to_Ioc_mod, to_Ioc_div_sub', sub_add_eq_add_sub] +by simp_rw [to_Ioc_mod, to_Ioc_div_sub', sub_right_comm] lemma to_Ico_mod_add_right' (a : α) {b : α} (hb : 0 < b) (x y : α) : to_Ico_mod a hb (x + y) = to_Ico_mod (a - y) hb x + y := -by simp_rw [to_Ico_mod, to_Ico_div_add_right', add_right_comm] +by simp_rw [to_Ico_mod, to_Ico_div_add_right', sub_add_eq_add_sub] lemma to_Ioc_mod_add_right' (a : α) {b : α} (hb : 0 < b) (x y : α) : to_Ioc_mod a hb (x + y) = to_Ioc_mod (a - y) hb x + y := -by simp_rw [to_Ioc_mod, to_Ioc_div_add_right', add_right_comm] +by simp_rw [to_Ioc_mod, to_Ioc_div_add_right', sub_add_eq_add_sub] lemma to_Ico_mod_neg (a : α) {b : α} (hb : 0 < b) (x : α) : to_Ico_mod a hb (-x) = b - to_Ioc_mod (-a) hb x := begin - simp_rw [to_Ico_mod, to_Ioc_mod, to_Ico_div_neg, sub_smul, one_smul], + simp_rw [to_Ico_mod, to_Ioc_mod, to_Ico_div_neg, neg_smul, add_smul], abel, end lemma to_Ioc_mod_neg (a : α) {b : α} (hb : 0 < b) (x : α) : to_Ioc_mod a hb (-x) = b - to_Ico_mod (-a) hb x := begin - simp_rw [to_Ioc_mod, to_Ico_mod, to_Ioc_div_neg, sub_smul, one_smul], + simp_rw [to_Ioc_mod, to_Ico_mod, to_Ioc_div_neg, neg_smul, add_smul], abel, end lemma to_Ico_mod_eq_to_Ico_mod (a : α) {b x y : α} (hb : 0 < b) : to_Ico_mod a hb x = to_Ico_mod a hb y ↔ ∃ z : ℤ, y - x = z • b := begin - refine ⟨λ h, ⟨to_Ico_div a hb x - to_Ico_div a hb y, _⟩, λ h, _⟩, - { conv_lhs { rw [←to_Ico_mod_sub_to_Ico_div_zsmul a hb x, - ←to_Ico_mod_sub_to_Ico_div_zsmul a hb y] }, + refine ⟨λ h, ⟨to_Ico_div a hb y - to_Ico_div a hb x, _⟩, λ h, _⟩, + { conv_lhs { rw [←to_Ico_mod_add_to_Ico_div_zsmul a hb x, + ←to_Ico_mod_add_to_Ico_div_zsmul a hb y] }, rw [h, sub_smul], abel }, { rcases h with ⟨z, hz⟩, @@ -446,9 +444,9 @@ end lemma to_Ioc_mod_eq_to_Ioc_mod (a : α) {b x y : α} (hb : 0 < b) : to_Ioc_mod a hb x = to_Ioc_mod a hb y ↔ ∃ z : ℤ, y - x = z • b := begin - refine ⟨λ h, ⟨to_Ioc_div a hb x - to_Ioc_div a hb y, _⟩, λ h, _⟩, - { conv_lhs { rw [←to_Ioc_mod_sub_to_Ioc_div_zsmul a hb x, - ←to_Ioc_mod_sub_to_Ioc_div_zsmul a hb y] }, + refine ⟨λ h, ⟨to_Ioc_div a hb y - to_Ioc_div a hb x, _⟩, λ h, _⟩, + { conv_lhs { rw [←to_Ioc_mod_add_to_Ioc_div_zsmul a hb x, + ←to_Ioc_mod_add_to_Ioc_div_zsmul a hb y] }, rw [h, sub_smul], abel }, { rcases h with ⟨z, hz⟩, @@ -464,7 +462,7 @@ omit hα /-- `mem_Ioo_mod a b x` means that `x` lies in the open interval `(a, a + b)` modulo `b`. Equivalently (as shown below), `x` is not congruent to `a` modulo `b`, or `to_Ico_mod a hb` agrees with `to_Ioc_mod a hb` at `x`, or `to_Ico_div a hb` agrees with `to_Ioc_div a hb` at `x`. -/ -def mem_Ioo_mod (b x : α) : Prop := ∃ z : ℤ, x + z • b ∈ set.Ioo a (a + b) +def mem_Ioo_mod (b x : α) : Prop := ∃ z : ℤ, x - z • b ∈ set.Ioo a (a + b) include hα lemma tfae_mem_Ioo_mod : @@ -475,15 +473,16 @@ lemma tfae_mem_Ioo_mod : begin tfae_have : 1 → 2, { exact λ ⟨i, hi⟩, - ((to_Ico_mod_eq_iff hb).2 ⟨set.Ioo_subset_Ico_self hi, i, add_sub_cancel' x _⟩).trans - ((to_Ioc_mod_eq_iff hb).2 ⟨set.Ioo_subset_Ioc_self hi, i, add_sub_cancel' x _⟩).symm }, + ((to_Ico_mod_eq_iff hb).2 ⟨set.Ioo_subset_Ico_self hi, i, (sub_add_cancel x _).symm⟩).trans + ((to_Ioc_mod_eq_iff hb).2 ⟨set.Ioo_subset_Ioc_self hi, i, (sub_add_cancel x _).symm⟩).symm }, tfae_have : 2 → 3, { intro h, rw [h, ne, add_right_eq_self], exact hb.ne' }, tfae_have : 3 → 4, { refine mt (λ h, _), rw [h, eq_comm, to_Ioc_mod_eq_iff, set.right_mem_Ioc], - refine ⟨lt_add_of_pos_right a hb, to_Ico_div a hb x + 1, _⟩, - conv_lhs { rw [← h, to_Ico_mod, add_assoc, ← add_one_zsmul, add_sub_cancel'] } }, + refine ⟨lt_add_of_pos_right a hb, to_Ico_div a hb x - 1, _⟩, + rw [sub_one_zsmul, add_add_add_comm, add_right_neg, add_zero], + conv_lhs { rw [← to_Ico_mod_add_to_Ico_div_zsmul a hb x, h] } }, tfae_have : 4 → 1, { have h' := to_Ico_mod_mem_Ico a hb x, exact λ h, ⟨_, h'.1.lt_of_ne' h, h'.2⟩ }, tfae_finish, @@ -501,43 +500,42 @@ lemma mem_Ioo_mod_iff_to_Ioc_mod_ne_right : mem_Ioo_mod a b x ↔ to_Ioc_mod a h begin rw [mem_Ioo_mod_iff_to_Ico_mod_eq_to_Ioc_mod, to_Ico_mod_eq_iff hb], obtain ⟨h₁, h₂⟩ := to_Ioc_mod_mem_Ioc a hb x, - exact ⟨λ h, h.1.2.ne, λ h, ⟨⟨h₁.le, h₂.lt_of_ne h⟩, _, add_sub_cancel' x _⟩⟩, + exact ⟨λ h, h.1.2.ne, λ h, ⟨⟨h₁.le, h₂.lt_of_ne h⟩, _, + (to_Ioc_mod_add_to_Ioc_div_zsmul _ _ _).symm⟩⟩, end lemma mem_Ioo_mod_iff_to_Ico_div_eq_to_Ioc_div : mem_Ioo_mod a b x ↔ to_Ico_div a hb x = to_Ioc_div a hb x := by rw [mem_Ioo_mod_iff_to_Ico_mod_eq_to_Ioc_mod hb, - to_Ico_mod, to_Ioc_mod, add_right_inj, (zsmul_strict_mono_left hb).injective.eq_iff] + to_Ico_mod, to_Ioc_mod, sub_right_inj, (zsmul_strict_mono_left hb).injective.eq_iff] lemma mem_Ioo_mod_iff_to_Ico_div_add_one_ne_to_Ioc_div : - mem_Ioo_mod a b x ↔ to_Ico_div a hb x + 1 ≠ to_Ioc_div a hb x := + mem_Ioo_mod a b x ↔ to_Ico_div a hb x ≠ to_Ioc_div a hb x + 1 := by rw [mem_Ioo_mod_iff_to_Ico_mod_add_period_ne_to_Ioc_mod hb, ne, ne, to_Ico_mod, to_Ioc_mod, - add_assoc, add_right_inj, ← add_one_zsmul, (zsmul_strict_mono_left hb).injective.eq_iff] + ← eq_sub_iff_add_eq, sub_sub, sub_right_inj, ← add_one_zsmul, + (zsmul_strict_mono_left hb).injective.eq_iff] include hb -lemma mem_Ioo_mod_iff_sub_ne_zsmul : mem_Ioo_mod a b x ↔ ∀ z : ℤ, a - x ≠ z • b := +lemma mem_Ioo_mod_iff_sub_ne_zsmul : mem_Ioo_mod a b x ↔ ∀ z : ℤ, x ≠ a + z • b := begin rw [mem_Ioo_mod_iff_to_Ico_mod_ne_left hb, ← not_iff_not], push_neg, split; intro h, { rw ← h, - exact ⟨_, add_sub_cancel' x _⟩ }, + exact ⟨_, (to_Ico_mod_add_to_Ico_div_zsmul _ _ _).symm⟩ }, { rw [to_Ico_mod_eq_iff, set.left_mem_Ico], exact ⟨lt_add_of_pos_right a hb, h⟩, }, end lemma mem_Ioo_mod_iff_eq_mod_zmultiples : mem_Ioo_mod a b x ↔ (x : α ⧸ add_subgroup.zmultiples b) ≠ a := -begin - rw [mem_Ioo_mod_iff_sub_ne_zsmul hb, ne, eq_comm, - quotient_add_group.eq_iff_sub_mem, add_subgroup.mem_zmultiples_iff], - push_neg, simp_rw ne_comm, -end +by simp_rw [mem_Ioo_mod_iff_sub_ne_zsmul hb, ne, quotient_add_group.eq_iff_sub_mem, + add_subgroup.mem_zmultiples_iff, eq_sub_iff_add_eq', eq_comm, not_exists] lemma Ico_eq_locus_Ioc_eq_Union_Ioo : - {x | to_Ico_mod a hb x = to_Ioc_mod a hb x} = ⋃ z : ℤ, set.Ioo (a - z • b) (a + b - z • b) := + {x | to_Ico_mod a hb x = to_Ioc_mod a hb x} = ⋃ z : ℤ, set.Ioo (a + z • b) (a + b + z • b) := begin - ext1, simp_rw [set.mem_set_of, set.mem_Union, ← set.add_mem_Ioo_iff_left], + ext1, simp_rw [set.mem_set_of, set.mem_Union, ← set.sub_mem_Ioo_iff_left], exact (mem_Ioo_mod_iff_to_Ico_mod_eq_to_Ioc_mod hb).symm, end @@ -561,28 +559,28 @@ end to_Ico_mod a₁ hb (to_Ico_mod a₂ hb x) = to_Ico_mod a₁ hb x := begin rw to_Ico_mod_eq_to_Ico_mod, - exact ⟨-to_Ico_div a₂ hb x, self_sub_to_Ico_mod a₂ hb x⟩ + exact ⟨to_Ico_div a₂ hb x, self_sub_to_Ico_mod a₂ hb x⟩ end @[simp] lemma to_Ico_mod_to_Ioc_mod (a₁ a₂ : α) {b : α} (hb : 0 < b) (x : α) : to_Ico_mod a₁ hb (to_Ioc_mod a₂ hb x) = to_Ico_mod a₁ hb x := begin rw to_Ico_mod_eq_to_Ico_mod, - exact ⟨-to_Ioc_div a₂ hb x, self_sub_to_Ioc_mod a₂ hb x⟩ + exact ⟨to_Ioc_div a₂ hb x, self_sub_to_Ioc_mod a₂ hb x⟩ end @[simp] lemma to_Ioc_mod_to_Ioc_mod (a₁ a₂ : α) {b : α} (hb : 0 < b) (x : α) : to_Ioc_mod a₁ hb (to_Ioc_mod a₂ hb x) = to_Ioc_mod a₁ hb x := begin rw to_Ioc_mod_eq_to_Ioc_mod, - exact ⟨-to_Ioc_div a₂ hb x, self_sub_to_Ioc_mod a₂ hb x⟩ + exact ⟨to_Ioc_div a₂ hb x, self_sub_to_Ioc_mod a₂ hb x⟩ end @[simp] lemma to_Ioc_mod_to_Ico_mod (a₁ a₂ : α) {b : α} (hb : 0 < b) (x : α) : to_Ioc_mod a₁ hb (to_Ico_mod a₂ hb x) = to_Ioc_mod a₁ hb x := begin rw to_Ioc_mod_eq_to_Ioc_mod, - exact ⟨-to_Ico_div a₂ hb x, self_sub_to_Ico_mod a₂ hb x⟩ + exact ⟨to_Ico_div a₂ hb x, self_sub_to_Ico_mod a₂ hb x⟩ end lemma to_Ico_mod_periodic (a : α) {b : α} (hb : 0 < b) : function.periodic (to_Ico_mod a hb) b := @@ -637,35 +635,34 @@ section linear_ordered_field variables {α : Type*} [linear_ordered_field α] [floor_ring α] -lemma to_Ico_div_eq_neg_floor (a : α) {b : α} (hb : 0 < b) (x : α) : - to_Ico_div a hb x = -⌊(x - a) / b⌋ := +lemma to_Ico_div_eq_floor (a : α) {b : α} (hb : 0 < b) (x : α) : + to_Ico_div a hb x = ⌊(x - a) / b⌋ := begin - refine (eq_to_Ico_div_of_add_zsmul_mem_Ico hb _).symm, - rw [set.mem_Ico, zsmul_eq_mul, int.cast_neg, neg_mul, ←sub_nonneg, add_comm, add_sub_assoc, - add_comm, ←sub_eq_add_neg], - refine ⟨int.sub_floor_div_mul_nonneg _ hb, _⟩, - rw [add_comm a, ←sub_lt_iff_lt_add, add_sub_assoc, add_comm, ←sub_eq_add_neg], - exact int.sub_floor_div_mul_lt _ hb + refine (eq_to_Ico_div_of_sub_zsmul_mem_Ico hb _).symm, + rw [set.mem_Ico, zsmul_eq_mul, ←sub_nonneg, add_comm, sub_right_comm, ←sub_lt_iff_lt_add, + sub_right_comm _ _ a], + exact ⟨int.sub_floor_div_mul_nonneg _ hb, int.sub_floor_div_mul_lt _ hb⟩, end -lemma to_Ioc_div_eq_floor (a : α) {b : α} (hb : 0 < b) (x : α) : - to_Ioc_div a hb x = ⌊(a + b - x) / b⌋ := +lemma to_Ioc_div_eq_neg_floor (a : α) {b : α} (hb : 0 < b) (x : α) : + to_Ioc_div a hb x = -⌊(a + b - x) / b⌋ := begin - refine (eq_to_Ioc_div_of_add_zsmul_mem_Ioc hb _).symm, - rw [set.mem_Ioc, zsmul_eq_mul, ←sub_nonneg, sub_add_eq_sub_sub], + refine (eq_to_Ioc_div_of_sub_zsmul_mem_Ioc hb _).symm, + rw [set.mem_Ioc, zsmul_eq_mul, int.cast_neg, neg_mul, sub_neg_eq_add, ←sub_nonneg, + sub_add_eq_sub_sub], refine ⟨_, int.sub_floor_div_mul_nonneg _ hb⟩, rw [←add_lt_add_iff_right b, add_assoc, add_comm x, ←sub_lt_iff_lt_add, add_comm (_ * _), ←sub_lt_iff_lt_add], exact int.sub_floor_div_mul_lt _ hb end -lemma to_Ico_div_zero_one (x : α) : to_Ico_div (0 : α) zero_lt_one x = -⌊x⌋ := -by simp [to_Ico_div_eq_neg_floor] +lemma to_Ico_div_zero_one (x : α) : to_Ico_div (0 : α) zero_lt_one x = ⌊x⌋ := +by simp [to_Ico_div_eq_floor] lemma to_Ico_mod_eq_add_fract_mul (a : α) {b : α} (hb : 0 < b) (x : α) : to_Ico_mod a hb x = a + int.fract ((x - a) / b) * b := begin - rw [to_Ico_mod, to_Ico_div_eq_neg_floor, int.fract], + rw [to_Ico_mod, to_Ico_div_eq_floor, int.fract], field_simp [hb.ne.symm], ring end @@ -677,7 +674,7 @@ by simp [to_Ico_mod_eq_add_fract_mul] lemma to_Ioc_mod_eq_sub_fract_mul (a : α) {b : α} (hb : 0 < b) (x : α) : to_Ioc_mod a hb x = a + b - int.fract ((a + b - x) / b) * b := begin - rw [to_Ioc_mod, to_Ioc_div_eq_floor, int.fract], + rw [to_Ioc_mod, to_Ioc_div_eq_neg_floor, int.fract], field_simp [hb.ne.symm], ring end diff --git a/src/analysis/special_functions/complex/arg.lean b/src/analysis/special_functions/complex/arg.lean index 59b2e80c434ae..07fb982444ab2 100644 --- a/src/analysis/special_functions/complex/arg.lean +++ b/src/analysis/special_functions/complex/arg.lean @@ -395,7 +395,7 @@ begin { convert to_Ioc_mod_mem_Ioc _ real.two_pi_pos _, ring }, convert arg_mul_cos_add_sin_mul_I hr hi using 3, - simp [to_Ioc_mod, cos_add_int_mul_two_pi, sin_add_int_mul_two_pi] + simp [to_Ioc_mod, cos_sub_int_mul_two_pi, sin_sub_int_mul_two_pi] end lemma arg_cos_add_sin_mul_I_eq_to_Ioc_mod (θ : ℝ) : @@ -405,7 +405,7 @@ by rw [←one_mul (_ + _), ←of_real_one, arg_mul_cos_add_sin_mul_I_eq_to_Ioc_m lemma arg_mul_cos_add_sin_mul_I_sub {r : ℝ} (hr : 0 < r) (θ : ℝ) : arg (r * (cos θ + sin θ * I)) - θ = 2 * π * ⌊(π - θ) / (2 * π)⌋ := begin - rw [arg_mul_cos_add_sin_mul_I_eq_to_Ioc_mod hr, to_Ioc_mod_sub_self, to_Ioc_div_eq_floor, + rw [arg_mul_cos_add_sin_mul_I_eq_to_Ioc_mod hr, to_Ioc_mod_sub_self, to_Ioc_div_eq_neg_floor, zsmul_eq_mul], ring_nf end diff --git a/src/analysis/special_functions/trigonometric/angle.lean b/src/analysis/special_functions/trigonometric/angle.lean index d670c850c3f76..d4e36aee7c0f4 100644 --- a/src/analysis/special_functions/trigonometric/angle.lean +++ b/src/analysis/special_functions/trigonometric/angle.lean @@ -407,14 +407,14 @@ end @[simp] lemma coe_to_Ico_mod (θ ψ : ℝ) : ↑(to_Ico_mod ψ two_pi_pos θ) = (θ : angle) := begin rw angle_eq_iff_two_pi_dvd_sub, - refine ⟨to_Ico_div ψ two_pi_pos θ, _⟩, + refine ⟨-to_Ico_div ψ two_pi_pos θ, _⟩, rw [to_Ico_mod_sub_self, zsmul_eq_mul, mul_comm] end @[simp] lemma coe_to_Ioc_mod (θ ψ : ℝ) : ↑(to_Ioc_mod ψ two_pi_pos θ) = (θ : angle) := begin rw angle_eq_iff_two_pi_dvd_sub, - refine ⟨to_Ioc_div ψ two_pi_pos θ, _⟩, + refine ⟨-to_Ioc_div ψ two_pi_pos θ, _⟩, rw [to_Ioc_mod_sub_self, zsmul_eq_mul, mul_comm] end diff --git a/src/topology/instances/add_circle.lean b/src/topology/instances/add_circle.lean index a67ddcb2eb982..98f0214adc573 100644 --- a/src/topology/instances/add_circle.lean +++ b/src/topology/instances/add_circle.lean @@ -69,12 +69,12 @@ begin let d := to_Ico_div a hp x • p, have hd := to_Ico_mod_mem_Ico a hp x, simp_rw [subset_def, mem_inter_iff], - refine ⟨_, ⟨l - d, min (a + p) u - d, _, λ x, id⟩, λ y, _⟩; - simp_rw [← add_mem_Ioo_iff_left, mem_Ioo, lt_min_iff], + refine ⟨_, ⟨l + d, min (a + p) u + d, _, λ x, id⟩, λ y, _⟩; + simp_rw [← sub_mem_Ioo_iff_left, mem_Ioo, lt_min_iff], { exact ⟨hxI.1, hd.2, hxI.2⟩ }, { rintro ⟨h, h'⟩, apply hIs, - rw [← to_Ico_mod_add_zsmul, (to_Ico_mod_eq_self _).2], - exacts [⟨h.1, h.2.2⟩, ⟨hd.1.trans (add_le_add_right h' _), h.2.1⟩] }, + rw [← to_Ico_mod_sub_zsmul, (to_Ico_mod_eq_self _).2], + exacts [⟨h.1, h.2.2⟩, ⟨hd.1.trans (sub_le_sub_right h' _), h.2.1⟩] }, end lemma continuous_left_to_Ioc_mod : continuous_within_at (to_Ioc_mod a hp) (Iic x) x := From a1b3559280c8e09220759127b400d1b02d7ac071 Mon Sep 17 00:00:00 2001 From: Riccardo Brasca Date: Sat, 14 Jan 2023 12:14:28 +0000 Subject: [PATCH 0244/1291] chore(algebra/order/ring/canonical): use implicit variables (#18167) Make implicitness of variables for `eq_zero_or_eq_zero_of_mul_eq_zero` consistent in `canonically_ordered_comm_semiring` and `no_zero_divisors_`. See the discussion in https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/Data.2ESet.2ESemiring Corresponds to [#1545](https://github.com/leanprover-community/mathlib4/pull/1545). --- src/algebra/order/ring/canonical.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/algebra/order/ring/canonical.lean b/src/algebra/order/ring/canonical.lean index 9727b9fcf4a59..1b4a386f746cd 100644 --- a/src/algebra/order/ring/canonical.lean +++ b/src/algebra/order/ring/canonical.lean @@ -38,7 +38,7 @@ not the integers or other ordered groups. -/ @[protect_proj, ancestor canonically_ordered_add_monoid comm_semiring] class canonically_ordered_comm_semiring (α : Type*) extends canonically_ordered_add_monoid α, comm_semiring α := -(eq_zero_or_eq_zero_of_mul_eq_zero : ∀ a b : α, a * b = 0 → a = 0 ∨ b = 0) +(eq_zero_or_eq_zero_of_mul_eq_zero : ∀ {a b : α}, a * b = 0 → a = 0 ∨ b = 0) section strict_ordered_semiring variables [strict_ordered_semiring α] {a b c d : α} @@ -81,7 +81,7 @@ variables [canonically_ordered_comm_semiring α] {a b : α} @[priority 100] -- see Note [lower instance priority] instance to_no_zero_divisors : no_zero_divisors α := -⟨canonically_ordered_comm_semiring.eq_zero_or_eq_zero_of_mul_eq_zero⟩ +⟨λ a b h, canonically_ordered_comm_semiring.eq_zero_or_eq_zero_of_mul_eq_zero h⟩ @[priority 100] -- see Note [lower instance priority] instance to_covariant_mul_le : covariant_class α α (*) (≤) := From 7a030ab8eb5d99f05a891dccc49c5b5b90c947d3 Mon Sep 17 00:00:00 2001 From: Paul Lezeau <72892199+Paul-Lez@users.noreply.github.com> Date: Sat, 14 Jan 2023 15:15:21 +0000 Subject: [PATCH 0245/1291] feat(ring_theory/polynomial/gauss_lemma): Prove Gauss's Lemma for integrally closed rings (#18147) In this PR, we prove Gauss's lemma for integrally closed rings. See #18021 and #11523 for previous discussion on the topic. We also show that integrally closed domains are precisely the domains in which Gauss's lemma holds for monic polynomials. [Zulip discussion](https://leanprover.zulipchat.com/#narrow/stream/144837-PR-reviews/topic/.2318021.20generalizing.20theory.20of.20minpoly) Co-authored-by: Junyan Xu Co-authored-by: Paul Lezeau --- src/data/polynomial/splits.lean | 11 ++ src/ring_theory/integral_closure.lean | 4 + src/ring_theory/integrally_closed.lean | 74 +++++++++++- src/ring_theory/polynomial/content.lean | 12 +- .../polynomial/cyclotomic/basic.lean | 2 +- src/ring_theory/polynomial/gauss_lemma.lean | 107 ++++++++++++++---- 6 files changed, 177 insertions(+), 33 deletions(-) diff --git a/src/data/polynomial/splits.lean b/src/data/polynomial/splits.lean index 107a88897b602..a28f5272e3995 100644 --- a/src/data/polynomial/splits.lean +++ b/src/data/polynomial/splits.lean @@ -5,6 +5,7 @@ Authors: Chris Hughes -/ import data.list.prime import data.polynomial.field_division +import data.polynomial.lifts /-! # Split polynomials @@ -311,6 +312,16 @@ begin simp, end +theorem mem_lift_of_splits_of_roots_mem_range (R : Type*) [comm_ring R] [algebra R K] {f : K[X]} + (hs : f.splits (ring_hom.id K)) (hm : f.monic) + (hr : ∀ a ∈ f.roots, a ∈ (algebra_map R K).range) : f ∈ polynomial.lifts (algebra_map R K) := +begin + rw [eq_prod_roots_of_monic_of_splits_id hm hs, lifts_iff_lifts_ring], + refine subring.multiset_prod_mem _ _ (λ P hP, _), + obtain ⟨b, hb, rfl⟩ := multiset.mem_map.1 hP, + exact subring.sub_mem _ (X_mem_lifts _) (C'_mem_lifts (hr _ hb)) +end + section UFD local attribute [instance, priority 10] principal_ideal_ring.to_unique_factorization_monoid diff --git a/src/ring_theory/integral_closure.lean b/src/ring_theory/integral_closure.lean index 230d8cbd0aa29..ffc6cb0bbd557 100644 --- a/src/ring_theory/integral_closure.lean +++ b/src/ring_theory/integral_closure.lean @@ -1031,4 +1031,8 @@ variables {R S : Type*} [comm_ring R] [comm_ring S] [is_domain S] [algebra R S] instance : is_domain (integral_closure R S) := infer_instance +theorem roots_mem_integral_closure {f : R[X]} (hf : f.monic) {a : S} + (ha : a ∈ (f.map $ algebra_map R S).roots) : a ∈ integral_closure R S := +⟨f, hf, (eval₂_eq_eval_map _).trans $ (mem_roots $ (hf.map _).ne_zero).1 ha⟩ + end is_domain diff --git a/src/ring_theory/integrally_closed.lean b/src/ring_theory/integrally_closed.lean index c22a258256162..a2483b48c5dc0 100644 --- a/src/ring_theory/integrally_closed.lean +++ b/src/ring_theory/integrally_closed.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Anne Baanen. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Anne Baanen -/ +import field_theory.splitting_field import ring_theory.integral_closure import ring_theory.localization.integral @@ -20,9 +21,13 @@ integral over `R`. A special case of integrally closed domains are the Dedekind * `is_integrally_closed_iff K`, where `K` is a fraction field of `R`, states `R` is integrally closed iff it is the integral closure of `R` in `K` +* `eq_map_mul_C_of_dvd`: if `K = Frac(R)` and `g : K[X]` divides a monic polynomial with + coefficients in `R`, then `g * (C g.leading_coeff⁻¹)` has coefficients in `R` -/ -open_locale non_zero_divisors +open_locale non_zero_divisors polynomial + +open polynomial /-- `R` is integrally closed if all integral elements of `Frac(R)` are also elements of `R`. @@ -124,8 +129,33 @@ namespace integral_closure open is_integrally_closed -variables {R : Type*} [comm_ring R] [is_domain R] -variables (K : Type*) [field K] [algebra R K] [is_fraction_ring R K] +variables {R : Type*} [comm_ring R] +variables (K : Type*) [field K] [algebra R K] + +theorem mem_lifts_of_monic_of_dvd_map + {f : R[X]} (hf : f.monic) {g : K[X]} (hg : g.monic) (hd : g ∣ f.map (algebra_map R K)) : + g ∈ lifts (algebra_map (integral_closure R K) K) := +begin + haveI : is_scalar_tower R K g.splitting_field := splitting_field_aux.is_scalar_tower _ _ _, + have := mem_lift_of_splits_of_roots_mem_range (integral_closure R g.splitting_field) + ((splits_id_iff_splits _).2 $ splitting_field.splits g) (hg.map _) + (λ a ha, (set_like.ext_iff.mp (integral_closure R g.splitting_field).range_algebra_map _).mpr $ + roots_mem_integral_closure hf _), + { rw [lifts_iff_coeff_lifts, ←ring_hom.coe_range, subalgebra.range_algebra_map] at this, + refine (lifts_iff_coeff_lifts _).2 (λ n, _), + rw [← ring_hom.coe_range, subalgebra.range_algebra_map], + obtain ⟨p, hp, he⟩ := (set_like.mem_coe.mp (this n)), use [p, hp], + rw [is_scalar_tower.algebra_map_eq R K, coeff_map, ← eval₂_map, eval₂_at_apply] at he, + rw eval₂_eq_eval_map, apply (injective_iff_map_eq_zero _).1 _ _ he, + { apply ring_hom.injective } }, + rw [is_scalar_tower.algebra_map_eq R K _, ← map_map], + refine multiset.mem_of_le (roots.le_of_dvd ((hf.map _).map _).ne_zero _) ha, + { apply_instance }, + { exact map_dvd (algebra_map K g.splitting_field) hd }, + { apply splitting_field_aux.is_scalar_tower }, +end + +variables [is_domain R] [is_fraction_ring R K] variables {L : Type*} [field L] [algebra K L] [algebra R L] [is_scalar_tower R K L] -- Can't be an instance because you need to supply `K`. @@ -137,3 +167,41 @@ begin end end integral_closure + +namespace is_integrally_closed + +open integral_closure + +variables {R : Type*} [comm_ring R] [is_domain R] +variables (K : Type*) [field K] [algebra R K] [is_fraction_ring R K] + +/-- If `K = Frac(R)` and `g : K[X]` divides a monic polynomial with coefficients in `R`, then + `g * (C g.leading_coeff⁻¹)` has coefficients in `R` -/ +lemma eq_map_mul_C_of_dvd [is_integrally_closed R] {f : R[X]} (hf : f.monic) + {g : K[X]} (hg : g ∣ f.map (algebra_map R K)) : + ∃ g' : R[X], (g'.map (algebra_map R K)) * (C $ leading_coeff g) = g := +begin + have g_ne_0 : g ≠ 0 := ne_zero_of_dvd_ne_zero (monic.ne_zero $ hf.map (algebra_map R K)) hg, + suffices lem : ∃ g' : R[X], g'.map (algebra_map R K) = g * (C g.leading_coeff⁻¹), + { obtain ⟨g', hg'⟩ := lem, + use g', + rw [hg', mul_assoc, ← C_mul, inv_mul_cancel (leading_coeff_ne_zero.mpr g_ne_0), C_1, mul_one] }, + + have g_mul_dvd : g * (C g.leading_coeff⁻¹) ∣ f.map (algebra_map R K), + { rwa associated.dvd_iff_dvd_left (show associated (g * (C (g.leading_coeff⁻¹))) g, from _), + rw associated_mul_is_unit_left_iff, + exact is_unit_C.mpr (inv_ne_zero $ leading_coeff_ne_zero.mpr g_ne_0).is_unit }, + let algeq := (subalgebra.equiv_of_eq _ _ $ integral_closure_eq_bot R _).trans + (algebra.bot_equiv_of_injective $ is_fraction_ring.injective R $ K), + have : (algebra_map R _).comp algeq.to_alg_hom.to_ring_hom = + (integral_closure R _).to_subring.subtype, + { ext, conv_rhs { rw ← algeq.symm_apply_apply x }, refl }, + have H := ((mem_lifts _ ).1 + (mem_lifts_of_monic_of_dvd_map K hf (monic_mul_leading_coeff_inv g_ne_0) g_mul_dvd)), + refine ⟨map algeq.to_alg_hom.to_ring_hom _, _⟩, + use classical.some H, + rw [map_map, this], + exact classical.some_spec H +end + +end is_integrally_closed diff --git a/src/ring_theory/polynomial/content.lean b/src/ring_theory/polynomial/content.lean index 0a614943c463c..89250e33b1380 100644 --- a/src/ring_theory/polynomial/content.lean +++ b/src/ring_theory/polynomial/content.lean @@ -58,6 +58,9 @@ begin exact (hp 0 (dvd_zero (C 0))).ne_zero rfl, end +lemma is_primitive_of_dvd {p q : R[X]} (hp : is_primitive p) (hq : q ∣ p) : is_primitive q := +λ a ha, is_primitive_iff_is_unit_of_C_dvd.mp hp a (dvd_trans ha hq) + end primitive variables {R : Type*} [comm_ring R] [is_domain R] @@ -382,15 +385,6 @@ begin ring, end -lemma is_primitive.is_primitive_of_dvd {p q : R[X]} (hp : p.is_primitive) (hdvd : q ∣ p) : - q.is_primitive := -begin - rcases hdvd with ⟨r, rfl⟩, - rw [is_primitive_iff_content_eq_one, ← normalize_content, normalize_eq_one, is_unit_iff_dvd_one], - apply dvd.intro r.content, - rwa [is_primitive_iff_content_eq_one, content_mul] at hp, -end - lemma is_primitive.dvd_prim_part_iff_dvd {p q : R[X]} (hp : p.is_primitive) (hq : q ≠ 0) : p ∣ q.prim_part ↔ p ∣ q := diff --git a/src/ring_theory/polynomial/cyclotomic/basic.lean b/src/ring_theory/polynomial/cyclotomic/basic.lean index 7ec13bdc53eed..c983afb708475 100644 --- a/src/ring_theory/polynomial/cyclotomic/basic.lean +++ b/src/ring_theory/polynomial/cyclotomic/basic.lean @@ -840,7 +840,7 @@ end lemma cyclotomic.irreducible_rat {n : ℕ} (hpos : 0 < n) : irreducible (cyclotomic n ℚ) := begin rw [← map_cyclotomic_int], - exact (is_primitive.int.irreducible_iff_irreducible_map_cast (cyclotomic.is_primitive n ℤ)).1 + exact (is_primitive.irreducible_iff_irreducible_map_fraction_map (cyclotomic.is_primitive n ℤ)).1 (cyclotomic.irreducible hpos), end diff --git a/src/ring_theory/polynomial/gauss_lemma.lean b/src/ring_theory/polynomial/gauss_lemma.lean index c0681583f97d5..50fbc6b2bc3fb 100644 --- a/src/ring_theory/polynomial/gauss_lemma.lean +++ b/src/ring_theory/polynomial/gauss_lemma.lean @@ -4,7 +4,10 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Aaron Anderson -/ import ring_theory.int.basic +import field_theory.splitting_field import ring_theory.localization.integral +import ring_theory.integrally_closed + /-! # Gauss's Lemma @@ -12,12 +15,18 @@ import ring_theory.localization.integral Gauss's Lemma is one of a few results pertaining to irreducibility of primitive polynomials. ## Main Results + - `polynomial.monic.irreducible_iff_irreducible_map_fraction_map`: + A monic polynomial over an integrally closed domain is irreducible iff it is irreducible in a + fraction field + - `is_integrally_closed_iff'`: + Integrally closed domains are precisely the domains for in which Gauss's lemma holds + for monic polynomials - `polynomial.is_primitive.irreducible_iff_irreducible_map_fraction_map`: - A primitive polynomial is irreducible iff it is irreducible in a fraction field. + A primitive polynomial over a GCD domain is irreducible iff it is irreducible in a fraction field - `polynomial.is_primitive.int.irreducible_iff_irreducible_map_cast`: A primitive polynomial over `ℤ` is irreducible iff it is irreducible over `ℚ`. - `polynomial.is_primitive.dvd_iff_fraction_map_dvd_fraction_map`: - Two primitive polynomials divide each other iff they do in a fraction field. + Two primitive polynomials over a GCD domain divide each other iff they do in a fraction field. - `polynomial.is_primitive.int.dvd_iff_map_cast_dvd_map_cast`: Two primitive polynomials over `ℤ` divide each other if they do in `ℚ`. @@ -25,15 +34,13 @@ Gauss's Lemma is one of a few results pertaining to irreducibility of primitive open_locale non_zero_divisors polynomial -variables {R : Type*} [comm_ring R] [is_domain R] +variables {R : Type*} [comm_ring R] namespace polynomial -section normalized_gcd_monoid -variable [normalized_gcd_monoid R] section -variables {S : Type*} [comm_ring S] [is_domain S] {φ : R →+* S} (hinj : function.injective φ) -variables {f : R[X]} (hf : f.is_primitive) +variables {S : Type*} [comm_ring S] [is_domain S] +variables {φ : R →+* S} (hinj : function.injective φ) {f : R[X]} (hf : f.is_primitive) include hinj hf lemma is_primitive.is_unit_iff_is_unit_map_of_injective : @@ -43,34 +50,93 @@ begin rcases is_unit_iff.1 h with ⟨_, ⟨u, rfl⟩, hu⟩, have hdeg := degree_C u.ne_zero, rw [hu, degree_map_eq_of_injective hinj] at hdeg, - rw [eq_C_of_degree_eq_zero hdeg, is_primitive_iff_content_eq_one, - content_C, normalize_eq_one] at hf, - rwa [eq_C_of_degree_eq_zero hdeg, is_unit_C], + rw [eq_C_of_degree_eq_zero hdeg] at hf ⊢, + exact is_unit_C.mpr (is_primitive_iff_is_unit_of_C_dvd.mp hf (f.coeff 0) dvd_rfl), end lemma is_primitive.irreducible_of_irreducible_map_of_injective (h_irr : irreducible (map φ f)) : irreducible f := begin - refine ⟨λ h, h_irr.not_unit (is_unit.map (map_ring_hom φ) h), _⟩, - intros a b h, - rcases h_irr.is_unit_or_is_unit (by rw [h, polynomial.map_mul]) with hu | hu, - { left, - rwa (hf.is_primitive_of_dvd (dvd.intro _ h.symm)).is_unit_iff_is_unit_map_of_injective hinj }, - right, - rwa (hf.is_primitive_of_dvd (dvd.intro_left _ h.symm)).is_unit_iff_is_unit_map_of_injective hinj + refine ⟨λ h, h_irr.not_unit (is_unit.map (map_ring_hom φ) h), + λ a b h, (h_irr.is_unit_or_is_unit $ by rw [h, polynomial.map_mul]).imp _ _⟩, + all_goals { apply ((is_primitive_of_dvd hf _).is_unit_iff_is_unit_map_of_injective hinj).mpr }, + exacts [(dvd.intro _ h.symm), dvd.intro_left _ h.symm], end end section fraction_map + variables {K : Type*} [field K] [algebra R K] [is_fraction_ring R K] lemma is_primitive.is_unit_iff_is_unit_map {p : R[X]} (hp : p.is_primitive) : is_unit p ↔ is_unit (p.map (algebra_map R K)) := hp.is_unit_iff_is_unit_map_of_injective (is_fraction_ring.injective _ _) +variable [is_domain R] + +section is_integrally_closed + +open is_integrally_closed + +/-- **Gauss's Lemma** for integrally closed domains states that a monic polynomial is irreducible + iff it is irreducible in the fraction field. -/ +theorem monic.irreducible_iff_irreducible_map_fraction_map [is_integrally_closed R] {p : R[X]} + (h : p.monic) : irreducible p ↔ irreducible (p.map $ algebra_map R K) := +begin + /- The ← direction follows from `is_primitive.irreducible_of_irreducible_map_of_injective`. + For the → direction, it is enought to show that if `(p.map $ algebra_map R K) = a * b` and + `a` is not a unit then `b` is a unit -/ + refine ⟨λ hp, irreducible_iff.mpr ⟨hp.not_unit.imp h.is_primitive.is_unit_iff_is_unit_map.mpr, + λ a b H, or_iff_not_imp_left.mpr (λ hₐ, _)⟩, + λ hp, h.is_primitive.irreducible_of_irreducible_map_of_injective + (is_fraction_ring.injective R K) hp⟩, + + obtain ⟨a', ha⟩ := eq_map_mul_C_of_dvd K h (dvd_of_mul_right_eq b H.symm), + obtain ⟨b', hb⟩ := eq_map_mul_C_of_dvd K h (dvd_of_mul_left_eq a H.symm), + + have : a.leading_coeff * b.leading_coeff = 1, + { rw [← leading_coeff_mul, ← H, monic.leading_coeff (h.map $ algebra_map R K)] }, + + rw [← ha, ← hb, mul_comm _ (C b.leading_coeff), mul_assoc, ← mul_assoc (C a.leading_coeff), + ← C_mul, this, C_1, one_mul, ← polynomial.map_mul] at H, + rw [← hb, ← polynomial.coe_map_ring_hom], + refine is_unit.mul + (is_unit.map _ (or.resolve_left (hp.is_unit_or_is_unit _) (show ¬ is_unit a', from _))) + (is_unit_iff_exists_inv'.mpr (exists.intro (C a.leading_coeff) $ by rwa [← C_mul, this, C_1])), + { exact polynomial.map_injective _ (is_fraction_ring.injective R K) H }, + + { by_contra h_contra, + refine hₐ _, + rw [← ha, ← polynomial.coe_map_ring_hom], + exact is_unit.mul (is_unit.map _ h_contra) (is_unit_iff_exists_inv.mpr + (exists.intro (C b.leading_coeff) $ by rwa [← C_mul, this, C_1])) }, +end + +/-- Integrally closed domains are precisely the domains for in which Gauss's lemma holds + for monic polynomials -/ +theorem is_integrally_closed_iff' : is_integrally_closed R ↔ + ∀ p : R[X], p.monic → (irreducible p ↔ irreducible (p.map $ algebra_map R K)) := +begin + split, + { intros hR p hp, letI := hR, exact monic.irreducible_iff_irreducible_map_fraction_map hp }, + { intro H, + refine (is_integrally_closed_iff K).mpr (λ x hx, ring_hom.mem_range.mp $ + minpoly.mem_range_of_degree_eq_one R x _), + rw ← monic.degree_map (minpoly.monic hx) (algebra_map R K), + apply degree_eq_one_of_irreducible_of_root ((H _ $ minpoly.monic hx).mp + (minpoly.irreducible hx)), + rw [is_root, eval_map, ← aeval_def, minpoly.aeval R x] }, +end + +end is_integrally_closed + open is_localization +section normalized_gcd_monoid + +variable [normalized_gcd_monoid R] + lemma is_unit_or_eq_zero_of_is_unit_integer_normalization_prim_part {p : K[X]} (h0 : p ≠ 0) (h : is_unit (integer_normalization R⁰ p).prim_part) : is_unit p := @@ -91,8 +157,8 @@ begin { apply units.ne_zero _ con }, end -/-- **Gauss's Lemma** states that a primitive polynomial is irreducible iff it is irreducible in the - fraction field. -/ +/-- **Gauss's Lemma** for GCD domains states that a primitive polynomial is irreducible iff it is + irreducible in the fraction field. -/ theorem is_primitive.irreducible_iff_irreducible_map_fraction_map {p : R[X]} (hp : p.is_primitive) : irreducible p ↔ irreducible (p.map (algebra_map R K)) := @@ -167,6 +233,8 @@ lemma is_primitive.dvd_iff_fraction_map_dvd_fraction_map {p q : R[X]} ⟨λ ⟨a,b⟩, ⟨a.map (algebra_map R K), b.symm ▸ polynomial.map_mul (algebra_map R K)⟩, λ h, hp.dvd_of_fraction_map_dvd_fraction_map hq h⟩ +end normalized_gcd_monoid + end fraction_map /-- **Gauss's Lemma** for `ℤ` states that a primitive integer polynomial is irreducible iff it is @@ -181,5 +249,4 @@ lemma is_primitive.int.dvd_iff_map_cast_dvd_map_cast (p q : ℤ[X]) (p ∣ q) ↔ (p.map (int.cast_ring_hom ℚ) ∣ q.map (int.cast_ring_hom ℚ)) := hp.dvd_iff_fraction_map_dvd_fraction_map ℚ hq -end normalized_gcd_monoid end polynomial From 761f917097b30b3cb3151a731ea0fe6221ca8f40 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sat, 14 Jan 2023 17:40:46 +0000 Subject: [PATCH 0246/1291] chore(archive/100-theorems-list/30_ballot_problem): golf (#18126) Also add some supporting lemmas. --- .../100-theorems-list/30_ballot_problem.lean | 163 ++++-------------- src/data/list/basic.lean | 2 + src/data/list/perm.lean | 16 ++ src/logic/basic.lean | 7 +- 4 files changed, 57 insertions(+), 131 deletions(-) diff --git a/archive/100-theorems-list/30_ballot_problem.lean b/archive/100-theorems-list/30_ballot_problem.lean index f0d659d67a7aa..5e9a632acceba 100644 --- a/archive/100-theorems-list/30_ballot_problem.lean +++ b/archive/100-theorems-list/30_ballot_problem.lean @@ -66,73 +66,44 @@ This represents vote sequences where candidate `+1` receives `p` votes and candi def counted_sequence (p q : ℕ) : set (list ℤ) := {l | l.count 1 = p ∧ l.count (-1) = q ∧ ∀ x ∈ l, x = (1 : ℤ) ∨ x = -1} -@[simp] lemma counted_right_zero (p : ℕ) : counted_sequence p 0 = {list.repeat 1 p} := +/-- An alternative definition of `counted_sequence` that uses `list.perm`. -/ +lemma mem_counted_sequence_iff_perm {p q l} : + l ∈ counted_sequence p q ↔ l ~ list.repeat (1 : ℤ) p ++ list.repeat (-1) q := begin - ext l, - rw [counted_sequence, mem_singleton_iff], - split, - { rintro ⟨hl₀, hl₁, hl₂⟩, - rw list.eq_repeat, - have : ∀ x ∈ l, (1 : ℤ) = x, - { intros x hx, - obtain rfl | rfl := hl₂ x hx, - { refl }, - { exact false.elim (list.not_mem_of_count_eq_zero hl₁ hx) } }, - split, - { rwa ← list.count_eq_length.2 this }, - { exact λ x hx, (this x hx).symm } }, - { rintro rfl, - simp only [mem_set_of_eq, list.count_repeat_self, eq_self_iff_true, true_and], - refine ⟨list.count_eq_zero_of_not_mem _, λ x, _⟩; rw list.mem_repeat, - { norm_num }, - { rintro ⟨-, rfl⟩, - exact or.inl rfl } } + rw [list.perm_repeat_append_repeat], + { simp only [counted_sequence, list.subset_def, mem_set_of_eq, list.mem_cons_iff, + list.mem_singleton] }, + { norm_num1 } end +@[simp] lemma counted_right_zero (p : ℕ) : counted_sequence p 0 = {list.repeat 1 p} := +by { ext l, simp [mem_counted_sequence_iff_perm] } + @[simp] lemma counted_left_zero (q : ℕ) : counted_sequence 0 q = {list.repeat (-1) q} := -begin - ext l, - rw [counted_sequence, mem_singleton_iff], - split, - { rintro ⟨hl₀, hl₁, hl₂⟩, - rw list.eq_repeat, - have : ∀ x ∈ l, (-1 : ℤ) = x, - { intros x hx, - obtain rfl | rfl := hl₂ x hx, - { exact false.elim (list.not_mem_of_count_eq_zero hl₀ hx) }, - { refl } }, - split, - { rwa ← list.count_eq_length.2 this }, - { exact λ x hx, (this x hx).symm } }, - { rintro rfl, - simp only [mem_set_of_eq, list.count_repeat_self, eq_self_iff_true, true_and], - refine ⟨list.count_eq_zero_of_not_mem _, λ x, _⟩; rw list.mem_repeat, - { norm_num }, - { rintro ⟨-, rfl⟩, - exact or.inr rfl } } -end +by { ext l, simp [mem_counted_sequence_iff_perm] } + +lemma mem_of_mem_counted_sequence {p q} {l} (hl : l ∈ counted_sequence p q) {x : ℤ} (hx : x ∈ l) : + x = 1 ∨ x = -1 := +hl.2.2 x hx + +lemma length_of_mem_counted_sequence {p q} {l : list ℤ} (hl : l ∈ counted_sequence p q) : + l.length = p + q := +by simp [(mem_counted_sequence_iff_perm.1 hl).length_eq] + +lemma counted_eq_nil_iff {p q : ℕ} {l : list ℤ} (hl : l ∈ counted_sequence p q) : + l = [] ↔ p = 0 ∧ q = 0 := +list.length_eq_zero.symm.trans $ by simp [length_of_mem_counted_sequence hl] lemma counted_ne_nil_left {p q : ℕ} (hp : p ≠ 0) {l : list ℤ} (hl : l ∈ counted_sequence p q) : - l ≠ list.nil := -begin - obtain ⟨hl₀, hl₁, hl₂⟩ := hl, - rintro rfl, - rw list.count_nil at hl₀, - exact hp hl₀.symm, -end + l ≠ [] := +by simp [counted_eq_nil_iff hl, hp] -lemma counted_ne_nil_right {p q : ℕ} (hp : q ≠ 0) {l : list ℤ} (hl : l ∈ counted_sequence p q) : - l ≠ list.nil := -begin - obtain ⟨hl₀, hl₁, hl₂⟩ := hl, - rintro rfl, - rw list.count_nil at hl₁, - exact hp hl₁.symm, -end +lemma counted_ne_nil_right {p q : ℕ} (hq : q ≠ 0) {l : list ℤ} (hl : l ∈ counted_sequence p q) : + l ≠ [] := +by simp [counted_eq_nil_iff hl, hq] lemma counted_succ_succ (p q : ℕ) : counted_sequence (p + 1) (q + 1) = - (counted_sequence p (q + 1)).image (list.cons 1) ∪ - (counted_sequence (p + 1) q).image (list.cons (-1)) := + list.cons 1 '' counted_sequence p (q + 1) ∪ list.cons (-1) '' counted_sequence (p + 1) q := begin ext l, rw [counted_sequence, counted_sequence, counted_sequence], @@ -193,79 +164,13 @@ lemma counted_sequence_nonempty : ∀ (p q : ℕ), (counted_sequence p q).nonemp exact or.inl (counted_sequence_nonempty _ _), end -lemma sum_of_mem_counted_sequence : - ∀ {p q : ℕ} {l : list ℤ} (hl : l ∈ counted_sequence p q), l.sum = p - q -| 0 q l hl := - begin - rw [counted_left_zero, mem_singleton_iff] at hl, - simp [hl], - end -| p 0 l hl := - begin - rw [counted_right_zero, mem_singleton_iff] at hl, - simp [hl], - end -| (p + 1) (q + 1) l hl := - begin - simp only [counted_succ_succ, mem_union, mem_image] at hl, - rcases hl with (⟨l, hl, rfl⟩ | ⟨l, hl, rfl⟩), - { rw [list.sum_cons, sum_of_mem_counted_sequence hl], - push_cast, - ring }, - { rw [list.sum_cons, sum_of_mem_counted_sequence hl], - push_cast, - ring } - end - -lemma mem_of_mem_counted_sequence : - ∀ {p q : ℕ} {l} (hl : l ∈ counted_sequence p q) {x : ℤ} (hx : x ∈ l), x = 1 ∨ x = -1 -| 0 q l hl x hx := - begin - rw [counted_left_zero, mem_singleton_iff] at hl, - subst hl, - exact or.inr (list.eq_of_mem_repeat hx), - end -| p 0 l hl x hx := - begin - rw [counted_right_zero, mem_singleton_iff] at hl, - subst hl, - exact or.inl (list.eq_of_mem_repeat hx), - end -| (p + 1) (q + 1) l hl x hx := - begin - simp only [counted_succ_succ, mem_union, mem_image] at hl, - rcases hl with (⟨l, hl, rfl⟩ | ⟨l, hl, rfl⟩); - rcases hx with (rfl | hx), - { left, refl }, - { exact mem_of_mem_counted_sequence hl hx }, - { right, refl }, - { exact mem_of_mem_counted_sequence hl hx }, - end - -lemma length_of_mem_counted_sequence : - ∀ {p q : ℕ} {l : list ℤ} (hl : l ∈ counted_sequence p q), l.length = p + q -| 0 q l hl := - begin - rw [counted_left_zero, mem_singleton_iff] at hl, - simp [hl], - end -| p 0 l hl := - begin - rw [counted_right_zero, mem_singleton_iff] at hl, - simp [hl], - end -| (p + 1) (q + 1) l hl := - begin - simp only [counted_succ_succ, mem_union, mem_image] at hl, - rcases hl with (⟨l, hl, rfl⟩ | ⟨l, hl, rfl⟩), - { rw [list.length_cons, length_of_mem_counted_sequence hl, add_right_comm] }, - { rw [list.length_cons, length_of_mem_counted_sequence hl, ←add_assoc] } - end +lemma sum_of_mem_counted_sequence {p q} {l : list ℤ} (hl : l ∈ counted_sequence p q) : + l.sum = p - q := +by simp [(mem_counted_sequence_iff_perm.1 hl).sum_eq, sub_eq_add_neg] lemma disjoint_bits (p q : ℕ) : - disjoint - ((counted_sequence p (q + 1)).image (list.cons 1)) - ((counted_sequence (p + 1) q).image (list.cons (-1))) := + disjoint (list.cons 1 '' counted_sequence p (q + 1)) + (list.cons (-1) '' counted_sequence (p + 1) q) := begin simp_rw [disjoint_left, mem_image, not_exists, exists_imp_distrib], rintros _ _ ⟨_, rfl⟩ _ ⟨_, _, _⟩, diff --git a/src/data/list/basic.lean b/src/data/list/basic.lean index 051098553ce50..913dbfa253faf 100644 --- a/src/data/list/basic.lean +++ b/src/data/list/basic.lean @@ -318,6 +318,8 @@ iff.intro or_exists_of_exists_mem_cons /-! ### list subset -/ +instance : is_trans (list α) (⊆) := ⟨λ _ _ _, list.subset.trans⟩ + theorem subset_def {l₁ l₂ : list α} : l₁ ⊆ l₂ ↔ ∀ ⦃a : α⦄, a ∈ l₁ → a ∈ l₂ := iff.rfl theorem subset_append_of_subset_left (l l₁ l₂ : list α) : l ⊆ l₁ → l ⊆ l₁++l₂ := diff --git a/src/data/list/perm.lean b/src/data/list/perm.lean index af5455996a7de..5064dbc38934e 100644 --- a/src/data/list/perm.lean +++ b/src/data/list/perm.lean @@ -79,6 +79,12 @@ theorem perm.subset {l₁ l₂ : list α} (p : l₁ ~ l₂) : l₁ ⊆ l₂ := theorem perm.mem_iff {a : α} {l₁ l₂ : list α} (h : l₁ ~ l₂) : a ∈ l₁ ↔ a ∈ l₂ := iff.intro (λ m, h.subset m) (λ m, h.symm.subset m) +lemma perm.subset_congr_left {l₁ l₂ l₃ : list α} (h : l₁ ~ l₂) : l₁ ⊆ l₃ ↔ l₂ ⊆ l₃ := +⟨h.symm.subset.trans, h.subset.trans⟩ + +lemma perm.subset_congr_right {l₁ l₂ l₃ : list α} (h : l₁ ~ l₂) : l₃ ⊆ l₁ ↔ l₃ ⊆ l₂ := +⟨λ h', h'.trans h.subset, λ h', h'.trans h.symm.subset⟩ + theorem perm.append_right {l₁ l₂ : list α} (t₁ : list α) (p : l₁ ~ l₂) : l₁++t₁ ~ l₂++t₁ := perm.rec_on p (perm.refl ([] ++ t₁)) @@ -730,6 +736,16 @@ theorem perm_iff_count {l₁ l₂ : list α} : l₁ ~ l₂ ↔ ∀ a, count a l by_cases b = a; simp [h] at H ⊢; assumption } end⟩ +theorem perm_repeat_append_repeat {l : list α} {a b : α} {m n : ℕ} (h : a ≠ b) : + l ~ repeat a m ++ repeat b n ↔ count a l = m ∧ count b l = n ∧ l ⊆ [a, b] := +begin + rw [perm_iff_count, ← decidable.and_forall_ne a, ← decidable.and_forall_ne b], + suffices : l ⊆ [a, b] ↔ ∀ c, c ≠ b → c ≠ a → c ∉ l, + { simp [count_repeat, h, h.symm, this] { contextual := tt } }, + simp_rw [ne.def, ← and_imp, ← not_or_distrib, decidable.not_imp_not, subset_def, mem_cons_iff, + not_mem_nil, or_false, or_comm], +end + lemma subperm.cons_right {α : Type*} {l l' : list α} (x : α) (h : l <+~ l') : l <+~ x :: l' := h.trans (sublist_cons x l').subperm diff --git a/src/logic/basic.lean b/src/logic/basic.lean index bd88f41f2a834..79bafe7763ff2 100644 --- a/src/logic/basic.lean +++ b/src/logic/basic.lean @@ -1146,10 +1146,13 @@ by simp [and_comm] @[simp] theorem forall_eq' {a' : α} : (∀a, a' = a → p a) ↔ p a' := by simp [@eq_comm _ a'] -theorem and_forall_ne (a : α) : (p a ∧ ∀ b ≠ a, p b) ↔ ∀ b, p b := -by simp only [← @forall_eq _ p a, ← forall_and_distrib, ← or_imp_distrib, classical.em, +theorem decidable.and_forall_ne [decidable_eq α] (a : α) : (p a ∧ ∀ b ≠ a, p b) ↔ ∀ b, p b := +by simp only [← @forall_eq _ p a, ← forall_and_distrib, ← or_imp_distrib, decidable.em, forall_const] +theorem and_forall_ne (a : α) : (p a ∧ ∀ b ≠ a, p b) ↔ ∀ b, p b := +decidable.and_forall_ne a + -- this lemma is needed to simplify the output of `list.mem_cons_iff` @[simp] theorem forall_eq_or_imp {a' : α} : (∀ a, a = a' ∨ q a → p a) ↔ p a' ∧ ∀ a, q a → p a := by simp only [or_imp_distrib, forall_and_distrib, forall_eq] From 0a6c26ee8a47bd28d7b0bf45ad459eb266e3da48 Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Sat, 14 Jan 2023 21:04:59 +0000 Subject: [PATCH 0247/1291] feat(ring_theory/power_basis): minpoly_gen is always the minimal polynomial (#18117) + add `minpoly.unique'`, a characterization for being the minimal polynomial. + golf various proofs and remove some unnecessary typeclass assumptions. --- src/data/polynomial/algebra_map.lean | 10 +- src/field_theory/minpoly/basic.lean | 131 +++++++++----------- src/number_theory/cyclotomic/gal.lean | 7 +- src/ring_theory/adjoin/power_basis.lean | 2 +- src/ring_theory/adjoin_root.lean | 22 ++-- src/ring_theory/power_basis.lean | 158 ++++++++---------------- 6 files changed, 127 insertions(+), 203 deletions(-) diff --git a/src/data/polynomial/algebra_map.lean b/src/data/polynomial/algebra_map.lean index cd83ab6081dc1..858805215b536 100644 --- a/src/data/polynomial/algebra_map.lean +++ b/src/data/polynomial/algebra_map.lean @@ -259,9 +259,9 @@ by rw [← algebra.map_top, ← adjoin_X, alg_hom.map_adjoin, set.image_singleto variable {R} -section comm_semiring +section semiring -variables [comm_semiring S] {f : R →+* S} +variables [semiring S] {f : R →+* S} lemma aeval_eq_sum_range [algebra R S] {p : R[X]} (x : S) : aeval x p = ∑ i in finset.range (p.nat_degree + 1), p.coeff i • x ^ i := @@ -284,9 +284,13 @@ lemma is_root_of_aeval_algebra_map_eq_zero [algebra R S] {p : R[X]} {r : R} (hr : aeval (algebra_map R S r) p = 0) : p.is_root r := is_root_of_eval₂_map_eq_zero inj hr +end semiring + +section comm_semiring + section aeval_tower -variables [algebra S R] [algebra S A'] [algebra S B'] +variables [comm_semiring S] [algebra S R] [algebra S A'] [algebra S B'] /-- Version of `aeval` for defining algebra homs out of `R[X]` over a smaller base ring than `R`. -/ diff --git a/src/field_theory/minpoly/basic.lean b/src/field_theory/minpoly/basic.lean index ec940ccbf05d3..fd66cfcd422f6 100644 --- a/src/field_theory/minpoly/basic.lean +++ b/src/field_theory/minpoly/basic.lean @@ -34,7 +34,7 @@ For example, if `V` is a `𝕜`-vector space for some field `𝕜` and `f : V the minimal polynomial of `f` is `minpoly 𝕜 f`. -/ noncomputable def minpoly (x : B) : A[X] := -if hx : is_integral A x then well_founded.min degree_lt_wf _ hx else 0 +if hx : is_integral A x then degree_lt_wf.min _ hx else 0 end min_poly_def @@ -46,7 +46,7 @@ variables {x : B} /-- A minimal polynomial is monic. -/ lemma monic (hx : is_integral A x) : monic (minpoly A x) := -by { delta minpoly, rw dif_pos hx, exact (well_founded.min_mem degree_lt_wf _ hx).1 } +by { delta minpoly, rw dif_pos hx, exact (degree_lt_wf.min_mem _ hx).1 } /-- A minimal polynomial is nonzero. -/ lemma ne_zero [nontrivial A] (hx : is_integral A x) : minpoly A x ≠ 0 := @@ -61,7 +61,7 @@ variables (A x) @[simp] lemma aeval : aeval x (minpoly A x) = 0 := begin delta minpoly, split_ifs with hx, - { exact (well_founded.min_mem degree_lt_wf _ hx).2 }, + { exact (degree_lt_wf.min_mem _ hx).2 }, { exact aeval_zero _ } end @@ -108,10 +108,32 @@ lemma min {p : A[X]} (pmonic : p.monic) (hp : polynomial.aeval x p = 0) : degree (minpoly A x) ≤ degree p := begin delta minpoly, split_ifs with hx, - { exact le_of_not_lt (well_founded.not_lt_min degree_lt_wf _ hx ⟨pmonic, hp⟩) }, + { exact le_of_not_lt (degree_lt_wf.not_lt_min _ hx ⟨pmonic, hp⟩) }, { simp only [degree_zero, bot_le] } end +lemma unique' {p : A[X]} (hm : p.monic) (hp : polynomial.aeval x p = 0) + (hl : ∀ q : A[X], degree q < degree p → q = 0 ∨ polynomial.aeval x q ≠ 0) : + p = minpoly A x := +begin + nontriviality A, + have hx : is_integral A x := ⟨p, hm, hp⟩, + obtain h | h := hl _ ((minpoly A x).degree_mod_by_monic_lt hm), swap, + { exact (h $ (aeval_mod_by_monic_eq_self_of_root hm hp).trans $ aeval A x).elim }, + obtain ⟨r, hr⟩ := (dvd_iff_mod_by_monic_eq_zero hm).1 h, + rw hr, have hlead := congr_arg leading_coeff hr, + rw [mul_comm, leading_coeff_mul_monic hm, (monic hx).leading_coeff] at hlead, + have : nat_degree r ≤ 0, + { have hr0 : r ≠ 0 := by { rintro rfl, exact ne_zero hx (mul_zero p ▸ hr) }, + apply_fun nat_degree at hr, + rw hm.nat_degree_mul' hr0 at hr, + apply nat.le_of_add_le_add_left, + rw add_zero, + exact hr.symm.trans_le (nat_degree_le_nat_degree $ min A x hm hp) }, + rw [eq_C_of_nat_degree_le_zero this, ← nat.eq_zero_of_le_zero this, + ← leading_coeff, ← hlead, C_1, mul_one], +end + @[nontriviality] lemma subsingleton [subsingleton B] : minpoly A x = 1 := begin nontriviality A, @@ -130,11 +152,11 @@ variables [comm_ring A] section ring -variables [ring B] [algebra A B] [nontrivial B] +variables [ring B] [algebra A B] variables {x : B} /-- The degree of a minimal polynomial, as a natural number, is positive. -/ -lemma nat_degree_pos (hx : is_integral A x) : 0 < nat_degree (minpoly A x) := +lemma nat_degree_pos [nontrivial B] (hx : is_integral A x) : 0 < nat_degree (minpoly A x) := begin rw pos_iff_ne_zero, intro ndeg_eq_zero, @@ -145,7 +167,7 @@ begin end /-- The degree of a minimal polynomial is positive. -/ -lemma degree_pos (hx : is_integral A x) : 0 < degree (minpoly A x) := +lemma degree_pos [nontrivial B] (hx : is_integral A x) : 0 < degree (minpoly A x) := nat_degree_pos_iff_degree_pos.mp (nat_degree_pos hx) /-- If `B/A` is an injective ring extension, and `a` is an element of `A`, @@ -155,34 +177,20 @@ lemma eq_X_sub_C_of_algebra_map_inj minpoly A (algebra_map A B a) = X - C a := begin nontriviality A, - have hdegle : (minpoly A (algebra_map A B a)).nat_degree ≤ 1, - { apply with_bot.coe_le_coe.1, - rw [←degree_eq_nat_degree (ne_zero (@is_integral_algebra_map A B _ _ _ a)), - with_top.coe_one, ←degree_X_sub_C a], - refine min A (algebra_map A B a) (monic_X_sub_C a) _, - simp only [aeval_C, aeval_X, alg_hom.map_sub, sub_self] }, - have hdeg : (minpoly A (algebra_map A B a)).degree = 1, - { apply (degree_eq_iff_nat_degree_eq (ne_zero (@is_integral_algebra_map A B _ _ _ a))).2, - apply le_antisymm hdegle (nat_degree_pos (@is_integral_algebra_map A B _ _ _ a)) }, - have hrw := eq_X_add_C_of_degree_eq_one hdeg, - simp only [monic (@is_integral_algebra_map A B _ _ _ a), one_mul, - monic.leading_coeff, ring_hom.map_one] at hrw, - have h0 : (minpoly A (algebra_map A B a)).coeff 0 = -a, - { have hroot := aeval A (algebra_map A B a), - rw [hrw, add_comm] at hroot, - simp only [aeval_C, aeval_X, aeval_add] at hroot, - replace hroot := eq_neg_of_add_eq_zero_left hroot, - rw [←ring_hom.map_neg _ a] at hroot, - exact (hf hroot) }, - rw hrw, - simp only [h0, ring_hom.map_neg, sub_eq_add_neg], + refine (unique' A _ (monic_X_sub_C a) _ _).symm, + { rw [map_sub, aeval_C, aeval_X, sub_self] }, + simp_rw or_iff_not_imp_left, + intros q hl h0, + rw [← nat_degree_lt_nat_degree_iff h0, nat_degree_X_sub_C, nat.lt_one_iff] at hl, + rw eq_C_of_nat_degree_eq_zero hl at h0 ⊢, + rwa [aeval_C, map_ne_zero_iff _ hf, ← C_ne_zero], end end ring section is_domain -variables [is_domain A] [ring B] [algebra A B] +variables [ring B] [algebra A B] variables {x : B} /-- If `a` strictly divides the minimal polynomial of `x`, then `x` cannot be a root for `a`. -/ @@ -190,59 +198,30 @@ lemma aeval_ne_zero_of_dvd_not_unit_minpoly {a : A[X]} (hx : is_integral A x) (hamonic : a.monic) (hdvd : dvd_not_unit a (minpoly A x)) : polynomial.aeval x a ≠ 0 := begin - intro ha, - refine not_lt_of_ge (minpoly.min A x hamonic ha) _, - obtain ⟨hzeroa, b, hb_nunit, prod⟩ := hdvd, - have hbmonic : b.monic, - { rw monic.def, - have := monic hx, - rwa [monic.def, prod, leading_coeff_mul, monic.def.mp hamonic, one_mul] at this }, - have hzerob : b ≠ 0 := hbmonic.ne_zero, - have degbzero : 0 < b.nat_degree, - { apply nat.pos_of_ne_zero, - intro h, - have h₁ := eq_C_of_nat_degree_eq_zero h, - rw [←h, ←leading_coeff, monic.def.1 hbmonic, C_1] at h₁, - rw h₁ at hb_nunit, - have := is_unit_one, - contradiction }, - rw [prod, degree_mul, degree_eq_nat_degree hzeroa, degree_eq_nat_degree hzerob], - exact_mod_cast lt_add_of_pos_right _ degbzero, + refine λ ha, (min A x hamonic ha).not_lt (degree_lt_degree _), + obtain ⟨b, c, hu, he⟩ := hdvd, + have hcm := hamonic.of_mul_monic_left (he.subst $ monic hx), + rw [he, hamonic.nat_degree_mul hcm], + apply nat.lt_add_of_zero_lt_left _ _ (lt_of_not_le $ λ h, hu _), + rw [eq_C_of_nat_degree_le_zero h, ← nat.eq_zero_of_le_zero h, + ← leading_coeff, hcm.leading_coeff, C_1], + exact is_unit_one, end -variables [is_domain B] +variables [is_domain A] [is_domain B] /-- A minimal polynomial is irreducible. -/ lemma irreducible (hx : is_integral A x) : irreducible (minpoly A x) := begin - cases irreducible_or_factor (minpoly A x) (not_is_unit A x) with hirr hred, - { exact hirr }, - exfalso, - obtain ⟨a, b, ha_nunit, hb_nunit, hab_eq⟩ := hred, - have coeff_prod : a.leading_coeff * b.leading_coeff = 1, - { rw [←monic.def.1 (monic hx), ←hab_eq], - simp only [leading_coeff_mul] }, - have hamonic : (a * C b.leading_coeff).monic, - { rw monic.def, - simp only [coeff_prod, leading_coeff_mul, leading_coeff_C] }, - have hbmonic : (b * C a.leading_coeff).monic, - { rw [monic.def, mul_comm], - simp only [coeff_prod, leading_coeff_mul, leading_coeff_C] }, - have prod : minpoly A x = (a * C b.leading_coeff) * (b * C a.leading_coeff), - { symmetry, - calc a * C b.leading_coeff * (b * C a.leading_coeff) - = a * b * (C a.leading_coeff * C b.leading_coeff) : by ring - ... = a * b * (C (a.leading_coeff * b.leading_coeff)) : by simp only [ring_hom.map_mul] - ... = a * b : by rw [coeff_prod, C_1, mul_one] - ... = minpoly A x : hab_eq }, - have hzero := aeval A x, - rw [prod, aeval_mul, mul_eq_zero] at hzero, - cases hzero, - { refine aeval_ne_zero_of_dvd_not_unit_minpoly hx hamonic _ hzero, - exact ⟨hamonic.ne_zero, _, mt is_unit_of_mul_is_unit_left hb_nunit, prod⟩ }, - { refine aeval_ne_zero_of_dvd_not_unit_minpoly hx hbmonic _ hzero, - rw mul_comm at prod, - exact ⟨hbmonic.ne_zero, _, mt is_unit_of_mul_is_unit_left ha_nunit, prod⟩ }, + refine (irreducible_of_monic (monic hx) $ ne_one A x).2 (λ f g hf hg he, _), + rw [← hf.is_unit_iff, ← hg.is_unit_iff], + by_contra' h, + have heval := congr_arg (polynomial.aeval x) he, + rw [aeval A x, aeval_mul, mul_eq_zero] at heval, + cases heval, + { exact aeval_ne_zero_of_dvd_not_unit_minpoly hx hf ⟨hf.ne_zero, g, h.2, he.symm⟩ heval }, + { refine aeval_ne_zero_of_dvd_not_unit_minpoly hx hg ⟨hg.ne_zero, f, h.1, _⟩ heval, + rw [mul_comm, he] }, end end is_domain diff --git a/src/number_theory/cyclotomic/gal.lean b/src/number_theory/cyclotomic/gal.lean index 528db295c59e2..17ef9a05c2563 100644 --- a/src/number_theory/cyclotomic/gal.lean +++ b/src/number_theory/cyclotomic/gal.lean @@ -113,9 +113,8 @@ let hζ := zeta_spec n K L, end, right_inv := λ x, begin simp only [monoid_hom.to_fun_eq_coe], - generalize_proofs _ _ h, - have key := hζ.aut_to_pow_spec K ((hζ.power_basis K).equiv_of_minpoly - ((hμ x).power_basis K) h), + generalize_proofs _ h, + have key := hζ.aut_to_pow_spec K ((hζ.power_basis K).equiv_of_minpoly ((hμ x).power_basis K) h), have := (hζ.power_basis K).equiv_of_minpoly_gen ((hμ x).power_basis K) h, rw hζ.power_basis_gen K at this, rw [this, is_primitive_root.power_basis_gen] at key, @@ -142,7 +141,7 @@ let hζ := (zeta_spec n K L).eq_pow_of_pow_eq_one hμ.pow_eq_one n.pos in lemma from_zeta_aut_spec : from_zeta_aut hμ h (zeta n K L) = μ := begin simp_rw [from_zeta_aut, aut_equiv_pow_symm_apply], - generalize_proofs _ hζ h _ hμ _, + generalize_proofs hζ h _ hμ _, rw [←hζ.power_basis_gen K] {occs := occurrences.pos [4]}, rw [power_basis.equiv_of_minpoly_gen, hμ.power_basis_gen K], convert h.some_spec.some_spec, diff --git a/src/ring_theory/adjoin/power_basis.lean b/src/ring_theory/adjoin/power_basis.lean index e400d257dc74d..c91fc07156929 100644 --- a/src/ring_theory/adjoin/power_basis.lean +++ b/src/ring_theory/adjoin/power_basis.lean @@ -39,7 +39,7 @@ begin have minpoly_eq := minpoly.eq_of_algebra_map_eq hST hx' rfl, apply @basis.mk (fin (minpoly K x).nat_degree) _ (adjoin K {x}) (λ i, ⟨x, subset_adjoin (set.mem_singleton x)⟩ ^ (i : ℕ)), - { have := hx'.linear_independent_pow, + { have := linear_independent_pow _, rwa minpoly_eq at this }, { rintros ⟨y, hy⟩ _, have := hx'.mem_span_pow, diff --git a/src/ring_theory/adjoin_root.lean b/src/ring_theory/adjoin_root.lean index d3f088f947c3c..d0418ea5c80d7 100644 --- a/src/ring_theory/adjoin_root.lean +++ b/src/ring_theory/adjoin_root.lean @@ -328,15 +328,10 @@ basis.of_equiv_fun right_inv := λ x, funext $ λ i, begin nontriviality R, simp only [mod_by_monic_hom_mk], - rw [(mod_by_monic_eq_self_iff hg).mpr, finset_sum_coeff, finset.sum_eq_single i]; - try { simp only [coeff_monomial, eq_self_iff_true, if_true] }, - { intros j _ hj, exact if_neg (fin.coe_injective.ne hj) }, - { intros, have := finset.mem_univ i, contradiction }, - { refine (degree_sum_le _ _).trans_lt ((finset.sup_lt_iff _).mpr (λ j _, _)), - { exact bot_lt_iff_ne_bot.mpr (mt degree_eq_bot.mp hg.ne_zero) }, - { refine (degree_monomial_le _ _).trans_lt _, - rw [degree_eq_nat_degree hg.ne_zero, with_bot.coe_lt_coe], - exact j.2 } }, + rw [(mod_by_monic_eq_self_iff hg).mpr, finset_sum_coeff], + { simp_rw [coeff_monomial, fin.coe_eq_coe, finset.sum_ite_eq', if_pos (finset.mem_univ _)] }, + { simp_rw ← C_mul_X_pow_eq_monomial, + exact (degree_eq_nat_degree $ hg.ne_zero).symm ▸ degree_sum_fin_lt _ }, end} -- This was moved after the definition to prevent a timeout @@ -393,7 +388,7 @@ begin have minpoly_eq : minpoly K (root f) = f' := minpoly_root hf, apply @basis.mk _ _ _ (λ (i : fin f.nat_degree), (root f ^ i.val)), { rw [← deg_f', ← minpoly_eq], - exact (is_integral_root hf).linear_independent_pow }, + exact linear_independent_pow (root f) }, { rintros y -, rw [← deg_f', ← minpoly_eq], apply (is_integral_root hf).mem_span_pow, @@ -462,7 +457,7 @@ end minpoly section is_domain -variables [comm_ring R] [is_domain R] [comm_ring S] [is_domain S] [algebra R S] +variables [comm_ring R] [comm_ring S] [algebra R S] variables (g : R[X]) (pb : _root_.power_basis R S) /-- If `S` is an extension of `R` with power basis `pb` and `g` is a monic polynomial over `R` @@ -478,6 +473,7 @@ def equiv' (h₁ : aeval (root g) (minpoly R pb.gen) = 0) (h₂ : aeval pb.gen g inv_fun := pb.lift (root g) h₁, left_inv := λ x, induction_on g x $ λ f, by rw [lift_hom_mk, pb.lift_aeval, aeval_eq], right_inv := λ x, begin + nontriviality S, obtain ⟨f, hf, rfl⟩ := pb.exists_eq_aeval x, rw [pb.lift_aeval, aeval_eq, lift_hom_mk] end, @@ -508,7 +504,7 @@ def equiv (f : F[X]) (hf : f ≠ 0) : begin rw [power_basis_gen, minpoly_root hf, polynomial.map_mul, roots_mul, polynomial.map_C, roots_C, add_zero, equiv.refl_apply], - { rw ← polynomial.map_mul, exact map_monic_ne_zero (monic_mul_leading_coeff_inv hf) } + rw ← polynomial.map_mul, exact map_monic_ne_zero (monic_mul_leading_coeff_inv hf) end)) end field @@ -656,7 +652,7 @@ namespace power_basis open adjoin_root alg_equiv -variables [comm_ring R] [is_domain R] [comm_ring S] [is_domain S] [algebra R S] +variables [comm_ring R] [comm_ring S] [algebra R S] /-- Let `α` have minimal polynomial `f` over `R` and `I` be an ideal of `R`, then `R[α] / (I) = (R[x] / (f)) / pS = (R/p)[x] / (f mod p)`. -/ diff --git a/src/ring_theory/power_basis.lean b/src/ring_theory/power_basis.lean index a1f0fdba7af3c..e442b65f7bd2e 100644 --- a/src/ring_theory/power_basis.lean +++ b/src/ring_theory/power_basis.lean @@ -29,8 +29,7 @@ gives a `power_basis` structure generated by `x`. ## Implementation notes -Throughout this file, `R`, `S`, ... are `comm_ring`s, `A`, `B`, ... are -`comm_ring` with `is_domain`s and `K`, `L`, ... are `field`s. +Throughout this file, `R`, `S`, `A`, `B` ... are `comm_ring`s, and `K`, `L`, ... are `field`s. `S` is an `R`-algebra, `B` is an `A`-algebra, `L` is a `K`-algebra. ## Tags @@ -42,11 +41,9 @@ power basis, powerbasis open polynomial open_locale polynomial -variables {R S T : Type*} [comm_ring R] [comm_ring S] [comm_ring T] -variables [algebra R S] [algebra S T] [algebra R T] [is_scalar_tower R S T] -variables {A B : Type*} [comm_ring A] - [comm_ring B] [is_domain B] [algebra A B] -variables {K L : Type*} [field K] [field L] [algebra K L] +variables {R S T : Type*} [comm_ring R] [ring S] [algebra R S] +variables {A B : Type*} [comm_ring A] [comm_ring B] [is_domain B] [algebra A B] +variables {K : Type*} [field K] /-- `pb : power_basis R S` states that `1, pb.gen, ..., pb.gen ^ (pb.dim - 1)` is a basis for the `R`-algebra `S` (viewed as `R`-module). @@ -143,16 +140,11 @@ open_locale big_operators variable [algebra A S] -/-- `pb.minpoly_gen` is a minimal polynomial for `pb.gen`. - -If `A` is not a field, it might not necessarily be *the* minimal polynomial, -however `nat_degree_minpoly` shows its degree is indeed minimal. --/ +/-- `pb.minpoly_gen` is the minimal polynomial for `pb.gen`. -/ noncomputable def minpoly_gen (pb : power_basis A S) : A[X] := X ^ pb.dim - ∑ (i : fin pb.dim), C (pb.basis.repr (pb.gen ^ pb.dim) i) * X ^ (i : ℕ) -@[simp] lemma aeval_minpoly_gen (pb : power_basis A S) : aeval pb.gen (minpoly_gen pb) = 0 := begin simp_rw [minpoly_gen, alg_hom.map_sub, alg_hom.map_sum, alg_hom.map_mul, alg_hom.map_pow, @@ -162,27 +154,25 @@ begin simp only [pb.coe_basis, zero_smul, eq_self_iff_true, implies_true_iff] end -lemma dim_le_nat_degree_of_root (h : power_basis A S) {p : A[X]} - (ne_zero : p ≠ 0) (root : aeval h.gen p = 0) : - h.dim ≤ p.nat_degree := +lemma minpoly_gen_monic (pb : power_basis A S) : monic (minpoly_gen pb) := +begin + nontriviality A, + apply (monic_X_pow _).sub_of_left _, + rw degree_X_pow, + exact degree_sum_fin_lt _ +end + +lemma dim_le_nat_degree_of_root (pb : power_basis A S) {p : A[X]} + (ne_zero : p ≠ 0) (root : aeval pb.gen p = 0) : + pb.dim ≤ p.nat_degree := begin refine le_of_not_lt (λ hlt, ne_zero _), - let p_coeff : fin (h.dim) → A := λ i, p.coeff i, - suffices : ∀ i, p_coeff i = 0, - { ext i, - by_cases hi : i < h.dim, - { exact this ⟨i, hi⟩ }, - exact coeff_eq_zero_of_nat_degree_lt (lt_of_lt_of_le hlt (le_of_not_gt hi)) }, - intro i, - refine linear_independent_iff'.mp h.basis.linear_independent _ _ _ i (finset.mem_univ _), - rw aeval_eq_sum_range' hlt at root, - rw finset.sum_fin_eq_sum_range, - convert root, - ext i, - split_ifs with hi, - { simp_rw [coe_basis, p_coeff, fin.coe_mk] }, - { rw [coeff_eq_zero_of_nat_degree_lt (lt_of_lt_of_le hlt (le_of_not_gt hi)), - zero_smul] } + rw [p.as_sum_range' _ hlt, finset.sum_range], + refine fintype.sum_eq_zero _ (λ i, _), + simp_rw [aeval_eq_sum_range' hlt, finset.sum_range, ← pb.basis_eq_pow] at root, + have := fintype.linear_independent_iff.1 pb.basis.linear_independent _ root, + dsimp only at this, + rw [this, monomial_zero_right], end lemma dim_le_degree_of_root (h : power_basis A S) {p : A[X]} @@ -191,10 +181,7 @@ lemma dim_le_degree_of_root (h : power_basis A S) {p : A[X]} by { rw [degree_eq_nat_degree ne_zero, with_bot.coe_le_coe], exact h.dim_le_nat_degree_of_root ne_zero root } -variables [is_domain A] - -@[simp] -lemma degree_minpoly_gen (pb : power_basis A S) : +lemma degree_minpoly_gen [nontrivial A] (pb : power_basis A S) : degree (minpoly_gen pb) = pb.dim := begin unfold minpoly_gen, @@ -202,49 +189,43 @@ begin apply degree_sum_fin_lt end -@[simp] -lemma nat_degree_minpoly_gen (pb : power_basis A S) : +lemma nat_degree_minpoly_gen [nontrivial A] (pb : power_basis A S) : nat_degree (minpoly_gen pb) = pb.dim := nat_degree_eq_of_degree_eq_some pb.degree_minpoly_gen -lemma minpoly_gen_monic (pb : power_basis A S) : monic (minpoly_gen pb) := +@[simp] +lemma minpoly_gen_eq (pb : power_basis A S) : pb.minpoly_gen = minpoly A pb.gen := begin - apply (monic_X_pow _).sub_of_left _, - rw degree_X_pow, - exact degree_sum_fin_lt _ + nontriviality A, + refine minpoly.unique' A _ pb.minpoly_gen_monic + pb.aeval_minpoly_gen (λ q hq, or_iff_not_imp_left.2 $ λ hn0 h0, _), + exact (pb.dim_le_degree_of_root hn0 h0).not_lt (pb.degree_minpoly_gen ▸ hq), end lemma is_integral_gen (pb : power_basis A S) : is_integral A pb.gen := ⟨minpoly_gen pb, minpoly_gen_monic pb, aeval_minpoly_gen pb⟩ @[simp] -lemma nat_degree_minpoly (pb : power_basis A S) : - (minpoly A pb.gen).nat_degree = pb.dim := -begin - refine le_antisymm _ - (dim_le_nat_degree_of_root pb (minpoly.ne_zero pb.is_integral_gen) (minpoly.aeval _ _)), - rw ← nat_degree_minpoly_gen, - apply nat_degree_le_of_degree_le, - rw ← degree_eq_nat_degree (minpoly_gen_monic pb).ne_zero, - exact minpoly.min _ _ (minpoly_gen_monic pb) (aeval_minpoly_gen pb) -end +lemma degree_minpoly [nontrivial A] (pb : power_basis A S) : degree (minpoly A pb.gen) = pb.dim := +by rw [← minpoly_gen_eq, degree_minpoly_gen] @[simp] -lemma minpoly_gen_eq [algebra K S] (pb : power_basis K S) : - pb.minpoly_gen = minpoly K pb.gen := -minpoly.unique K pb.gen pb.minpoly_gen_monic pb.aeval_minpoly_gen (λ p p_monic p_root, - pb.degree_minpoly_gen.symm ▸ pb.dim_le_degree_of_root p_monic.ne_zero p_root) +lemma nat_degree_minpoly [nontrivial A] (pb : power_basis A S) : + (minpoly A pb.gen).nat_degree = pb.dim := +by rw [← minpoly_gen_eq, nat_degree_minpoly_gen] end minpoly section equiv -variables [algebra A S] {S' : Type*} [comm_ring S'] [algebra A S'] [is_domain A] +variables [algebra A S] {S' : Type*} [ring S'] [algebra A S'] lemma constr_pow_aeval (pb : power_basis A S) {y : S'} (hy : aeval y (minpoly A pb.gen) = 0) (f : A[X]) : pb.basis.constr A (λ i, y ^ (i : ℕ)) (aeval pb.gen f) = aeval y f := begin + casesI subsingleton_or_nontrivial A, + { rw [(subsingleton.elim _ _ : f = 0), aeval_zero, map_zero, aeval_zero] }, rw [← aeval_mod_by_monic_eq_self_of_root (minpoly.monic pb.is_integral_gen) (minpoly.aeval _ _), ← @aeval_mod_by_monic_eq_self_of_root _ _ _ _ _ f _ (minpoly.monic pb.is_integral_gen) y hy], by_cases hf : f %ₘ minpoly A pb.gen = 0, @@ -422,63 +403,30 @@ open power_basis /-- Useful lemma to show `x` generates a power basis: the powers of `x` less than the degree of `x`'s minimal polynomial are linearly independent. -/ -lemma is_integral.linear_independent_pow [algebra K S] {x : S} (hx : is_integral K x) : +lemma linear_independent_pow [algebra K S] (x : S) : linear_independent K (λ (i : fin (minpoly K x).nat_degree), x ^ (i : ℕ)) := begin - rw linear_independent_iff, - intros p hp, - set f : K[X] := p.sum (λ i, monomial i) with hf0, - have f_def : ∀ (i : fin _), f.coeff i = p i, - { intro i, - simp only [f, finsupp.sum, coeff_monomial, finset_sum_coeff], - rw [finset.sum_eq_single, if_pos rfl], - { intros b _ hb, - rw if_neg (mt (λ h, _) hb), - exact fin.coe_injective h }, - { intro hi, - rw if_pos rfl, - exact finsupp.not_mem_support_iff.mp hi } }, - have f_def' : ∀ i, f.coeff i = if hi : i < _ then p ⟨i, hi⟩ else 0, - { intro i, - split_ifs with hi, - { exact f_def ⟨i, hi⟩ }, - simp only [f, finsupp.sum, coeff_monomial, finset_sum_coeff], - apply finset.sum_eq_zero, - rintro ⟨j, hj⟩ -, - apply if_neg (mt _ hi), - rintro rfl, - exact hj }, - suffices : f = 0, - { ext i, rw [← f_def, this, coeff_zero, finsupp.zero_apply] }, - contrapose hp with hf, - intro h, - have : (minpoly K x).degree ≤ f.degree, - { apply minpoly.degree_le_of_ne_zero K x hf, - convert h, - simp_rw [finsupp.total_apply, aeval_def, hf0, finsupp.sum, eval₂_finset_sum], - apply finset.sum_congr rfl, - rintro i -, - simp only [algebra.smul_def, eval₂_monomial] }, - have : ¬ (minpoly K x).degree ≤ f.degree, - { apply not_le_of_lt, - rw [degree_eq_nat_degree (minpoly.ne_zero hx), degree_lt_iff_coeff_zero], - intros i hi, - rw [f_def' i, dif_neg], - exact hi.not_lt }, - contradiction + by_cases is_integral K x, swap, + { rw [minpoly.eq_zero h, nat_degree_zero], exact linear_independent_empty_type }, + refine fintype.linear_independent_iff.2 (λ g hg i, _), + simp only at hg, + simp_rw [algebra.smul_def, ← aeval_monomial, ← map_sum] at hg, + apply (λ hn0, (minpoly.degree_le_of_ne_zero K x (mt (λ h0, _) hn0) hg).not_lt).mtr, + { simp_rw ← C_mul_X_pow_eq_monomial, + exact (degree_eq_nat_degree $ minpoly.ne_zero h).symm ▸ degree_sum_fin_lt _ }, + { apply_fun lcoeff K i at h0, + simp_rw [map_sum, lcoeff_apply, coeff_monomial, fin.coe_eq_coe, finset.sum_ite_eq'] at h0, + exact (if_pos $ finset.mem_univ _).symm.trans h0 }, end lemma is_integral.mem_span_pow [nontrivial R] {x y : S} (hx : is_integral R x) (hy : ∃ f : R[X], y = aeval x f) : - y ∈ submodule.span R (set.range (λ (i : fin (minpoly R x).nat_degree), - x ^ (i : ℕ))) := + y ∈ submodule.span R (set.range (λ (i : fin (minpoly R x).nat_degree), x ^ (i : ℕ))) := begin obtain ⟨f, rfl⟩ := hy, apply mem_span_pow'.mpr _, have := minpoly.monic hx, - refine ⟨f.mod_by_monic (minpoly R x), - lt_of_lt_of_le (degree_mod_by_monic_lt _ this) degree_le_nat_degree, - _⟩, + refine ⟨f %ₘ minpoly R x, (degree_mod_by_monic_lt _ this).trans_le degree_le_nat_degree, _⟩, conv_lhs { rw ← mod_by_monic_add_div f this }, simp only [add_zero, zero_mul, minpoly.aeval, aeval_add, alg_hom.map_mul] end @@ -508,8 +456,6 @@ by { dsimp only [minpoly_gen, map_dim], -- Turn `fin (pb.map e).dim` into `fin p map_gen, alg_equiv.to_linear_equiv_apply, e.to_linear_equiv_symm, alg_equiv.map_pow, alg_equiv.symm_apply_apply, sub_right_inj] } -variables [is_domain A] - @[simp] lemma equiv_of_root_map (pb : power_basis A S) (e : S ≃ₐ[A] S') (h₁ h₂) : From dc17b6ac53b111affde68d96e5e7a0726816e2cf Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Sat, 14 Jan 2023 21:05:01 +0000 Subject: [PATCH 0248/1291] feat(algebra, linear_algebra): unique instances over the trivial ring (#18160) Also generalizes module.subsingleton to mul_action_with_zero. matches https://github.com/leanprover-community/mathlib4/pull/1519 --- src/algebra/module/basic.lean | 5 ++--- src/algebra/smul_with_zero.lean | 9 +++++++++ src/linear_algebra/basic.lean | 13 ++++++++++--- src/linear_algebra/basis.lean | 3 +++ src/linear_algebra/finite_dimensional.lean | 2 +- 5 files changed, 25 insertions(+), 7 deletions(-) diff --git a/src/algebra/module/basic.lean b/src/algebra/module/basic.lean index a242e541adc8a..14d70984fe3a3 100644 --- a/src/algebra/module/basic.lean +++ b/src/algebra/module/basic.lean @@ -263,14 +263,13 @@ as an instance because Lean has no way to guess `R`. -/ protected theorem module.subsingleton (R M : Type*) [semiring R] [subsingleton R] [add_comm_monoid M] [module R M] : subsingleton M := -⟨λ x y, by rw [← one_smul R x, ← one_smul R y, subsingleton.elim (1:R) 0, zero_smul, zero_smul]⟩ +mul_action_with_zero.subsingleton R M /-- A semiring is `nontrivial` provided that there exists a nontrivial module over this semiring. -/ protected theorem module.nontrivial (R M : Type*) [semiring R] [nontrivial M] [add_comm_monoid M] [module R M] : nontrivial R := -(subsingleton_or_nontrivial R).resolve_left $ λ hR, not_subsingleton M $ - by exactI module.subsingleton R M +mul_action_with_zero.nontrivial R M @[priority 910] -- see Note [lower instance priority] instance semiring.to_module [semiring R] : module R R := diff --git a/src/algebra/smul_with_zero.lean b/src/algebra/smul_with_zero.lean index e37fcce9ff05c..d8cb21c62c171 100644 --- a/src/algebra/smul_with_zero.lean +++ b/src/algebra/smul_with_zero.lean @@ -139,6 +139,15 @@ instance monoid_with_zero.to_opposite_mul_action_with_zero : mul_action_with_zer { ..mul_zero_class.to_opposite_smul_with_zero R, ..monoid.to_opposite_mul_action R } +protected lemma mul_action_with_zero.subsingleton + [mul_action_with_zero R M] [subsingleton R] : subsingleton M := +⟨λ x y, by rw [←one_smul R x, ←one_smul R y, subsingleton.elim (1 : R) 0, zero_smul, zero_smul]⟩ + +protected lemma mul_action_with_zero.nontrivial + [mul_action_with_zero R M] [nontrivial M] : nontrivial R := +(subsingleton_or_nontrivial R).resolve_left $ λ hR, not_subsingleton M $ + by exactI mul_action_with_zero.subsingleton R M + variables {R M} [mul_action_with_zero R M] [has_zero M'] [has_smul R M'] /-- Pullback a `mul_action_with_zero` structure along an injective zero-preserving homomorphism. diff --git a/src/linear_algebra/basic.lean b/src/linear_algebra/basic.lean index 379751730d4b5..ba3b23ccbe75f 100644 --- a/src/linear_algebra/basic.lean +++ b/src/linear_algebra/basic.lean @@ -1478,14 +1478,16 @@ namespace linear_equiv section add_comm_monoid section subsingleton -variables [semiring R] [semiring R₂] [semiring R₃] [semiring R₄] -variables [add_comm_monoid M] [add_comm_monoid M₂] [add_comm_monoid M₃] [add_comm_monoid M₄] +variables [semiring R] [semiring R₂] +variables [add_comm_monoid M] [add_comm_monoid M₂] variables [module R M] [module R₂ M₂] -variables [subsingleton M] [subsingleton M₂] variables {σ₁₂ : R →+* R₂} {σ₂₁ : R₂ →+* R} variables [ring_hom_inv_pair σ₁₂ σ₂₁] [ring_hom_inv_pair σ₂₁ σ₁₂] include σ₂₁ +section module +variables [subsingleton M] [subsingleton M₂] + /-- Between two zero modules, the zero map is an equivalence. -/ instance : has_zero (M ≃ₛₗ[σ₁₂] M₂) := ⟨{ to_fun := 0, @@ -1508,6 +1510,11 @@ instance : unique (M ≃ₛₗ[σ₁₂] M₂) := default := 0 } omit σ₂₁ +end module + +instance unique_of_subsingleton [subsingleton R] [subsingleton R₂] : unique (M ≃ₛₗ[σ₁₂] M₂) := +by { haveI := module.subsingleton R M, haveI := module.subsingleton R₂ M₂, apply_instance } + end subsingleton section diff --git a/src/linear_algebra/basis.lean b/src/linear_algebra/basis.lean index f5da0f90de12a..1fb2f9e25a40e 100644 --- a/src/linear_algebra/basis.lean +++ b/src/linear_algebra/basis.lean @@ -90,6 +90,9 @@ structure basis := of_repr :: (repr : M ≃ₗ[R] (ι →₀ R)) end +instance unique_basis [subsingleton R] : unique (basis ι R M) := +⟨⟨⟨default⟩⟩, λ ⟨b⟩, by rw subsingleton.elim b⟩ + namespace basis instance : inhabited (basis ι R (ι →₀ R)) := ⟨basis.of_repr (linear_equiv.refl _ _)⟩ diff --git a/src/linear_algebra/finite_dimensional.lean b/src/linear_algebra/finite_dimensional.lean index d0ee8c0ce5983..74d8af2c6886e 100644 --- a/src/linear_algebra/finite_dimensional.lean +++ b/src/linear_algebra/finite_dimensional.lean @@ -210,7 +210,7 @@ begin end /-- Given a basis of a division ring over itself indexed by a type `ι`, then `ι` is `unique`. -/ -noncomputable def basis.unique {ι : Type*} (b : basis ι K K) : unique ι := +noncomputable def _root_.basis.unique {ι : Type*} (b : basis ι K K) : unique ι := begin have A : cardinal.mk ι = ↑(finite_dimensional.finrank K K) := (finite_dimensional.finrank_eq_card_basis' b).symm, From dc0f25acd0b7a1f59becc08474d2901220266949 Mon Sep 17 00:00:00 2001 From: Michael Stoll Date: Sat, 14 Jan 2023 23:59:24 +0000 Subject: [PATCH 0249/1291] feat(number_theory/diophantine_approximation): add new file proving Dirichlet's approximation theorem (#18019) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR adds a new file, `number_theory/diophantine_approximation.lean`, that contains proofs of various versions of **Dirichlet's approximation theorem**, ```lean lemma real.exists_int_int_abs_mul_sub_le (ξ : ℝ) {n : ℕ} (n_pos : 0 < n) : ∃ j k : ℤ, 0 < k ∧ k ≤ n ∧ |↑k * ξ - j| ≤ 1 / (n + 1) lemma real.exists_nat_abs_mul_sub_round_le (ξ : ℝ) {n : ℕ} (n_pos : 0 < n) : ∃ k : ℕ, 0 < k ∧ k ≤ n ∧ |↑k * ξ - round (↑k * ξ)| ≤ 1 / (n + 1) ``` and ```lean lemma real.exists_rat_abs_sub_le_and_denom_le (ξ : ℝ) {n : ℕ} (n_pos : 0 < n) : ∃ q : ℚ, |ξ - q| ≤ 1 / ((n + 1) * q.denom) ∧ q.denom ≤ n ``` and also the following consequence: ```lean lemma real.rat_approx_infinite_of_irrational {ξ : ℝ} (hξ : irrational ξ) : {q : ℚ | |ξ - q| < 1 / q.denom ^ 2}.infinite ``` The proof of `exists_int_int_abs_mul_sub_le` is the standard one based on the pigeonhole principle. See also the [discussion on Zulip](https://leanprover.zulipchat.com/#narrow/stream/116395-maths/topic/Diophantine.20approximation/near/318002909). --- .../diophantine_approximation.lean | 181 ++++++++++++++++++ 1 file changed, 181 insertions(+) create mode 100644 src/number_theory/diophantine_approximation.lean diff --git a/src/number_theory/diophantine_approximation.lean b/src/number_theory/diophantine_approximation.lean new file mode 100644 index 0000000000000..61334a8a3aa0b --- /dev/null +++ b/src/number_theory/diophantine_approximation.lean @@ -0,0 +1,181 @@ +/- +Copyright (c) 2022 Michael Stoll. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Michael Geißer, Michael Stoll +-/ +import tactic.basic +import data.real.irrational +import combinatorics.pigeonhole + +/-! +# Diophantine Approximation + +This file gives proofs of various versions of **Dirichlet's approximation theorem** +and its important consequence that when `ξ` is an irrational real number, then there are +infinitely many rationals `x/y` (in lowest terms) such that `|ξ - x/y| < 1/y^2`. + +The proof is based on the pigeonhole principle. + +## Main statements + +The main results are three variants of Dirichlet's approximation theorem: +* `real.exists_int_int_abs_mul_sub_le`, which states that for all real `ξ` and natural `0 < n`, + there are integers `j` and `k` with `0 < k ≤ n` and `|k*ξ - j| ≤ 1/(n+1)`, +* `real.exists_nat_abs_mul_sub_round_le`, which replaces `j` by `round(k*ξ)` and uses + a natural number `k`, +* `real.exists_rat_abs_sub_le_and_denom_le`, which says that there is a rational number `q` + satisfying `|ξ - q| ≤ 1/((n+1)*q.denom)` and `q.denom ≤ n`, + +and +* `dioph_approx.rat_approx_infinite`, which states that for irrational `ξ`, the set + `{q : ℚ | |ξ - q| < 1/q.denom^2}` is infinite, + +## Implementation notes + +We use the namespace `real` for the results. + +## References + + + +## Tags + +Diophantine approximation, Dirichlet's approximation theorem +-/ + +namespace real + +section dirichlet + +/-! +### Dirichlet's approximation theorem + +We show that for any real number `ξ` and positive natural `n`, there is a fraction `q` +such that `q.denom ≤ n` and `|ξ - q| ≤ 1/((n+1)*q.denom)`. +-/ + +open finset int + +/-- *Dirichlet's approximation theorem:* +For any real number `ξ` and positive natural `n`, there are integers `j` and `k`, +with `0 < k ≤ n` and `|k*ξ - j| ≤ 1/(n+1)`. + +See also `real.exists_nat_abs_mul_sub_round_le`. -/ +lemma exists_int_int_abs_mul_sub_le (ξ : ℝ) {n : ℕ} (n_pos : 0 < n) : + ∃ j k : ℤ, 0 < k ∧ k ≤ n ∧ |↑k * ξ - j| ≤ 1 / (n + 1) := +begin + let f : ℤ → ℤ := λ m, ⌊fract (ξ * m) * (n + 1)⌋, + have hn : 0 < (n : ℝ) + 1 := by exact_mod_cast nat.succ_pos _, + have hfu := λ m : ℤ, mul_lt_of_lt_one_left hn $ fract_lt_one (ξ * ↑m), + conv in (|_| ≤ _) { rw [mul_comm, le_div_iff hn, ← abs_of_pos hn, ← abs_mul], }, + let D := Icc (0 : ℤ) n, + by_cases H : ∃ m ∈ D, f m = n, + { obtain ⟨m, hm, hf⟩ := H, + have hf' : ((n : ℤ) : ℝ) ≤ fract (ξ * m) * (n + 1) := hf ▸ floor_le (fract (ξ * m) * (n + 1)), + have hm₀ : 0 < m, + { have hf₀ : f 0 = 0, + { simp only [floor_eq_zero_iff, algebra_map.coe_zero, mul_zero, fract_zero, zero_mul, + set.left_mem_Ico, zero_lt_one], }, + refine ne.lt_of_le (λ h, n_pos.ne _) (mem_Icc.mp hm).1, + exact_mod_cast hf₀.symm.trans (h.symm ▸ hf : f 0 = n), }, + refine ⟨⌊ξ * m⌋ + 1, m, hm₀, (mem_Icc.mp hm).2, _⟩, + rw [cast_add, ← sub_sub, sub_mul, cast_one, one_mul, abs_le], + refine ⟨le_sub_iff_add_le.mpr _, + sub_le_iff_le_add.mpr $ le_of_lt $ (hfu m).trans $ lt_one_add _⟩, + simpa only [neg_add_cancel_comm_assoc] using hf', }, + { simp_rw [not_exists] at H, + have hD : (Ico (0 : ℤ) n).card < D.card, + { rw [card_Icc, card_Ico], exact lt_add_one n, }, + have hfu' : ∀ m, f m ≤ n := λ m, lt_add_one_iff.mp (floor_lt.mpr (by exact_mod_cast hfu m)), + have hwd : ∀ m : ℤ, m ∈ D → f m ∈ Ico (0 : ℤ) n := + λ x hx, mem_Ico.mpr ⟨floor_nonneg.mpr (mul_nonneg (fract_nonneg (ξ * x)) hn.le), + ne.lt_of_le (H x hx) (hfu' x)⟩, + have : ∃ (x : ℤ) (hx : x ∈ D) (y : ℤ) (hy : y ∈ D), x < y ∧ f x = f y, + { obtain ⟨x, hx, y, hy, x_ne_y, hxy⟩ := exists_ne_map_eq_of_card_lt_of_maps_to hD hwd, + rcases lt_trichotomy x y with h | h | h, + exacts [⟨x, hx, y, hy, h, hxy⟩, false.elim (x_ne_y h), ⟨y, hy, x, hx, h, hxy.symm⟩], }, + obtain ⟨x, hx, y, hy, x_lt_y, hxy⟩ := this, + refine ⟨⌊ξ * y⌋ - ⌊ξ * x⌋, y - x, sub_pos_of_lt x_lt_y, + sub_le_iff_le_add.mpr $ le_add_of_le_of_nonneg (mem_Icc.mp hy).2 (mem_Icc.mp hx).1, _⟩, + convert_to |fract (ξ * y) * (n + 1) - fract (ξ * x) * (n + 1)| ≤ 1, + { congr, push_cast, simp only [fract], ring, }, + exact (abs_sub_lt_one_of_floor_eq_floor hxy.symm).le, } +end + +/-- *Dirichlet's approximation theorem:* +For any real number `ξ` and positive natural `n`, there is a natural number `k`, +with `0 < k ≤ n` such that `|k*ξ - round(k*ξ)| ≤ 1/(n+1)`. +-/ +lemma exists_nat_abs_mul_sub_round_le (ξ : ℝ) {n : ℕ} (n_pos : 0 < n) : + ∃ k : ℕ, 0 < k ∧ k ≤ n ∧ |↑k * ξ - round (↑k * ξ)| ≤ 1 / (n + 1) := +begin + obtain ⟨j, k, hk₀, hk₁, h⟩ := exists_int_int_abs_mul_sub_le ξ n_pos, + have hk := to_nat_of_nonneg hk₀.le, + rw [← hk] at hk₀ hk₁ h, + exact ⟨k.to_nat, coe_nat_pos.mp hk₀, nat.cast_le.mp hk₁, (round_le (↑k.to_nat * ξ) j).trans h⟩, +end + +/-- *Dirichlet's approximation theorem:* +For any real number `ξ` and positive natural `n`, there is a fraction `q` +such that `q.denom ≤ n` and `|ξ - q| ≤ 1/((n+1)*q.denom)`. -/ +lemma exists_rat_abs_sub_le_and_denom_le (ξ : ℝ) {n : ℕ} (n_pos : 0 < n) : + ∃ q : ℚ, |ξ - q| ≤ 1 / ((n + 1) * q.denom) ∧ q.denom ≤ n := +begin + obtain ⟨j, k, hk₀, hk₁, h⟩ := exists_int_int_abs_mul_sub_le ξ n_pos, + have hk₀' : (0 : ℝ) < k := int.cast_pos.mpr hk₀, + have hden : ((j / k : ℚ).denom : ℤ) ≤ k, + { convert le_of_dvd hk₀ (rat.denom_dvd j k), exact rat.coe_int_div_eq_mk, }, + refine ⟨j / k, _, nat.cast_le.mp (hden.trans hk₁)⟩, + rw [← div_div, le_div_iff (nat.cast_pos.mpr $ rat.pos _ : (0 : ℝ) < _)], + refine (mul_le_mul_of_nonneg_left (int.cast_le.mpr hden : _ ≤ (k : ℝ)) (abs_nonneg _)).trans _, + rwa [← abs_of_pos hk₀', rat.cast_div, rat.cast_coe_int, rat.cast_coe_int, + ← abs_mul, sub_mul, div_mul_cancel _ hk₀'.ne', mul_comm], +end + +end dirichlet + +section rat_approx + +/-! +### Infinitely many good approximations to irrational numbers + +We show that an irrational real number `ξ` has infinitely many "good rational approximations", +i.e., fractions `x/y` in lowest terms such that `|ξ - x/y| < 1/y^2`. +-/ + +open set + +/-- Given any rational approximation `q` to the irrational real number `ξ`, there is +a good rational approximation `q'` such that `|ξ - q'| < |ξ - q|`. -/ +lemma exists_rat_abs_sub_lt_and_lt_of_irrational {ξ : ℝ} (hξ : irrational ξ) (q : ℚ) : + ∃ q' : ℚ, |ξ - q'| < 1 / q'.denom ^ 2 ∧ |ξ - q'| < |ξ - q| := +begin + have h := abs_pos.mpr (sub_ne_zero.mpr $ irrational.ne_rat hξ q), + obtain ⟨m, hm⟩ := exists_nat_gt (1 / |ξ - q|), + have m_pos : (0 : ℝ) < m := (one_div_pos.mpr h).trans hm, + obtain ⟨q', hbd, hden⟩ := exists_rat_abs_sub_le_and_denom_le ξ (nat.cast_pos.mp m_pos), + have den_pos : (0 : ℝ) < q'.denom := nat.cast_pos.mpr q'.pos, + have md_pos := mul_pos (add_pos m_pos zero_lt_one) den_pos, + refine ⟨q', lt_of_le_of_lt hbd _, + lt_of_le_of_lt hbd $ (one_div_lt md_pos h).mpr $ hm.trans $ + lt_of_lt_of_le (lt_add_one _) $ (le_mul_iff_one_le_right $ + add_pos m_pos zero_lt_one).mpr $ by exact_mod_cast (q'.pos : 1 ≤ q'.denom)⟩, + rw [sq, one_div_lt_one_div md_pos (mul_pos den_pos den_pos), mul_lt_mul_right den_pos], + exact lt_add_of_le_of_pos (nat.cast_le.mpr hden) zero_lt_one, +end + +/-- If `ξ` is an irrational real number, then there are infinitely many good +rational approximations to `ξ`. -/ +lemma infinite_rat_abs_sub_lt_one_div_denom_sq_of_irrational {ξ : ℝ} (hξ : irrational ξ) : + {q : ℚ | |ξ - q| < 1 / q.denom ^ 2}.infinite := +begin + refine or.resolve_left (set.finite_or_infinite _) (λ h, _), + obtain ⟨q, _, hq⟩ := exists_min_image {q : ℚ | |ξ - q| < 1 / q.denom ^ 2} (λ q, |ξ - q|) h + ⟨⌊ξ⌋, by simp [abs_of_nonneg, int.fract_lt_one]⟩, + obtain ⟨q', hmem, hbetter⟩ := exists_rat_abs_sub_lt_and_lt_of_irrational hξ q, + exact lt_irrefl _ (lt_of_le_of_lt (hq q' hmem) hbetter), +end + +end rat_approx + +end real From 88b76e4c78d85d9ac31d991aa05ff22c09da889b Mon Sep 17 00:00:00 2001 From: Michail Karatarakis Date: Sun, 15 Jan 2023 01:05:32 +0000 Subject: [PATCH 0250/1291] feat(ring_theory/valuation/valuation_subring): Add `valuation_subring.inertia_subgroup` (#17086) The decomposition and inertia subgroups. Co-authored-by: Eric Wieser Co-authored-by: Michail Karatarakis <40603357+mkaratarakis@users.noreply.github.com> --- src/ring_theory/ideal/local_ring.lean | 11 ++++ .../valuation/ramification_group.lean | 52 +++++++++++++++++++ 2 files changed, 63 insertions(+) create mode 100644 src/ring_theory/valuation/ramification_group.lean diff --git a/src/ring_theory/ideal/local_ring.lean b/src/ring_theory/ideal/local_ring.lean index 3f53ef4a90068..258f38590b2e2 100644 --- a/src/ring_theory/ideal/local_ring.lean +++ b/src/ring_theory/ideal/local_ring.lean @@ -389,6 +389,17 @@ is the residue field of `R`. -/ map_mul' := λ e₁ e₂, map_equiv_trans e₂ e₁, map_one' := map_equiv_refl } +section mul_semiring_action +variables (G : Type*) [group G] [mul_semiring_action G R] + +/-- If `G` acts on `R` as a `mul_semiring_action`, then it also acts on `residue_field R`. -/ +instance : mul_semiring_action G (local_ring.residue_field R) := +mul_semiring_action.comp_hom _ $ map_aut.comp (mul_semiring_action.to_ring_aut G R) + +@[simp] lemma residue_smul (g : G) (r : R) : residue R (g • r) = g • residue R r := rfl + +end mul_semiring_action + end residue_field lemma ker_eq_maximal_ideal [field K] (φ : R →+* K) (hφ : function.surjective φ) : diff --git a/src/ring_theory/valuation/ramification_group.lean b/src/ring_theory/valuation/ramification_group.lean new file mode 100644 index 0000000000000..9276b19d54fb1 --- /dev/null +++ b/src/ring_theory/valuation/ramification_group.lean @@ -0,0 +1,52 @@ +/- +Copyright (c) 2022 Michail Karatarakis. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Michail Karatarakis +-/ +import ring_theory.ideal.local_ring +import ring_theory.valuation.valuation_subring + +/-! +# Ramification groups + +The decomposition subgroup and inertia subgroups. + +TODO: Define higher ramification groups in lower numbering +-/ + +namespace valuation_subring + +open_locale pointwise + +variables (K : Type*) {L : Type*} [field K] [field L] [algebra K L] + +/-- The decomposition subgroup defined as the stabilizer of the action +on the type of all valuation subrings of the field. -/ +@[reducible] def decomposition_subgroup (A : valuation_subring L) : + subgroup (L ≃ₐ[K] L) := +mul_action.stabilizer (L ≃ₐ[K] L) A + +/-- The valuation subring `A` (considered as a subset of `L`) +is stable under the action of the decomposition group. -/ +def sub_mul_action (A : valuation_subring L) : + sub_mul_action (A.decomposition_subgroup K) L := +{ carrier := A, + smul_mem' := λ g l h, set.mem_of_mem_of_subset (set.smul_mem_smul_set h) g.prop.le } + +/-- The multiplicative action of the decomposition subgroup on `A`. -/ +instance decomposition_subgroup_mul_semiring_action (A : valuation_subring L) : + mul_semiring_action (A.decomposition_subgroup K) A := +{ smul_add := λ g k l, subtype.ext $ smul_add g k l, + smul_zero := λ g, subtype.ext $ smul_zero g, + smul_one := λ g, subtype.ext $ smul_one g, + smul_mul := λ g k l, subtype.ext $ smul_mul' g k l, + ..(sub_mul_action.mul_action (A.sub_mul_action K)) } + +/-- The inertia subgroup defined as the kernel of the group homomorphism from +the decomposition subgroup to the group of automorphisms of the residue field of `A`. -/ +def inertia_subgroup (A : valuation_subring L) : + subgroup (A.decomposition_subgroup K) := +monoid_hom.ker $ + mul_semiring_action.to_ring_aut (A.decomposition_subgroup K) (local_ring.residue_field A) + +end valuation_subring From f694c7dead66f5d4c80f446c796a5aad14707f0e Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sun, 15 Jan 2023 06:15:30 +0000 Subject: [PATCH 0251/1291] refactor(*): define `list.replicate` and migrate to it (#18127) This definition differs from `list.repeat` by the order of arguments. The new order is in sync with the Lean 4 version. --- .../100-theorems-list/30_ballot_problem.lean | 10 +- archive/100-theorems-list/45_partition.lean | 12 +- archive/miu_language/decision_suf.lean | 99 +++++---- src/algebra/big_operators/basic.lean | 6 +- src/algebra/big_operators/multiset/basic.lean | 10 +- src/algebra/gcd_monoid/multiset.lean | 4 +- src/algebra/hom/freiman.lean | 12 +- src/algebra/module/dedekind_domain.lean | 2 +- src/algebra/monoid_algebra/degree.lean | 4 +- src/algebra/polynomial/big_operators.lean | 4 +- src/combinatorics/composition.lean | 12 +- src/computability/turing_machine.lean | 34 +-- src/data/bitvec/core.lean | 8 +- src/data/dfinsupp/multiset.lean | 12 +- src/data/finset/basic.lean | 4 +- src/data/finset/sym.lean | 12 +- src/data/fintype/card.lean | 3 +- src/data/list/basic.lean | 151 ++++++------- src/data/list/big_operators/basic.lean | 16 +- src/data/list/big_operators/lemmas.lean | 2 +- src/data/list/count.lean | 26 +-- src/data/list/dedup.lean | 6 +- src/data/list/defs.lean | 5 + src/data/list/duplicate.lean | 2 +- src/data/list/nodup.lean | 6 +- src/data/list/nodup_equiv_fin.lean | 4 +- src/data/list/of_fn.lean | 2 +- src/data/list/pairwise.lean | 6 +- src/data/list/perm.lean | 32 ++- src/data/list/rotate.lean | 31 ++- src/data/list/sigma.lean | 2 +- src/data/list/sort.lean | 4 +- src/data/multiset/basic.lean | 208 ++++++++---------- src/data/multiset/bind.lean | 3 +- src/data/multiset/interval.lean | 2 +- src/data/multiset/nodup.lean | 2 +- src/data/nat/choose/multinomial.lean | 6 +- src/data/nat/digits.lean | 12 +- src/data/nat/factorization/basic.lean | 2 +- src/data/nat/factors.lean | 12 +- src/data/pnat/factors.lean | 6 +- src/data/polynomial/ring_division.lean | 9 +- src/data/polynomial/splits.lean | 2 +- .../qpf/multivariate/constructions/cofix.lean | 2 +- src/data/set/pointwise/list_of_fn.lean | 2 +- src/data/sym/basic.lean | 56 ++--- src/data/vector/basic.lean | 10 +- src/field_theory/abel_ruffini.lean | 2 +- src/group_theory/p_group.lean | 2 +- src/group_theory/perm/cycle/type.lean | 27 +-- src/group_theory/perm/sign.lean | 12 +- .../specific_groups/alternating.lean | 7 +- src/logic/equiv/list.lean | 4 +- src/number_theory/arithmetic_function.lean | 4 +- src/number_theory/ramification_inertia.lean | 4 +- src/ring_theory/chain_of_divisors.lean | 6 +- src/ring_theory/dedekind_domain/ideal.lean | 6 +- src/ring_theory/discrete_valuation_ring.lean | 22 +- src/ring_theory/polynomial/vieta.lean | 4 +- .../unique_factorization_domain.lean | 26 +-- src/set_theory/ordinal/fixed_point.lean | 2 +- src/tactic/explode.lean | 2 +- src/tactic/induction.lean | 2 +- src/tactic/omega/prove_unsats.lean | 10 +- src/tactic/solve_by_elim.lean | 2 +- src/tactic/unify_equations.lean | 2 +- 66 files changed, 502 insertions(+), 521 deletions(-) diff --git a/archive/100-theorems-list/30_ballot_problem.lean b/archive/100-theorems-list/30_ballot_problem.lean index 5e9a632acceba..1ba4b7e1f4299 100644 --- a/archive/100-theorems-list/30_ballot_problem.lean +++ b/archive/100-theorems-list/30_ballot_problem.lean @@ -68,18 +68,18 @@ def counted_sequence (p q : ℕ) : set (list ℤ) := /-- An alternative definition of `counted_sequence` that uses `list.perm`. -/ lemma mem_counted_sequence_iff_perm {p q l} : - l ∈ counted_sequence p q ↔ l ~ list.repeat (1 : ℤ) p ++ list.repeat (-1) q := + l ∈ counted_sequence p q ↔ l ~ list.replicate p (1 : ℤ) ++ list.replicate q (-1) := begin - rw [list.perm_repeat_append_repeat], + rw [list.perm_replicate_append_replicate], { simp only [counted_sequence, list.subset_def, mem_set_of_eq, list.mem_cons_iff, list.mem_singleton] }, { norm_num1 } end -@[simp] lemma counted_right_zero (p : ℕ) : counted_sequence p 0 = {list.repeat 1 p} := +@[simp] lemma counted_right_zero (p : ℕ) : counted_sequence p 0 = {list.replicate p 1} := by { ext l, simp [mem_counted_sequence_iff_perm] } -@[simp] lemma counted_left_zero (q : ℕ) : counted_sequence 0 q = {list.repeat (-1) q} := +@[simp] lemma counted_left_zero (q : ℕ) : counted_sequence 0 q = {list.replicate q (-1)} := by { ext l, simp [mem_counted_sequence_iff_perm] } lemma mem_of_mem_counted_sequence {p q} {l} (hl : l ∈ counted_sequence p q) {x : ℤ} (hx : x ∈ l) : @@ -286,7 +286,7 @@ begin rw mem_singleton_iff at hl, subst hl, refine λ l hl₁ hl₂, list.sum_pos _ (λ x hx, _) hl₁, - rw list.eq_of_mem_repeat (list.mem_of_mem_suffix hx hl₂), + rw list.eq_of_mem_replicate (list.mem_of_mem_suffix hx hl₂), norm_num }, end diff --git a/archive/100-theorems-list/45_partition.lean b/archive/100-theorems-list/45_partition.lean index 6c60da78735da..0fe42d3f7c6f1 100644 --- a/archive/100-theorems-list/45_partition.lean +++ b/archive/100-theorems-list/45_partition.lean @@ -335,19 +335,19 @@ begin { rwa [nat.nsmul_eq_mul, nat.nsmul_eq_mul, mul_left_inj' i.succ_ne_zero] at h } }, { simp only [mem_filter, mem_cut, mem_univ, exists_prop, true_and, and_assoc], rintros f ⟨hf₁, hf₂, hf₃⟩, - refine ⟨⟨∑ i in s, multiset.repeat i (f i / i), _, _⟩, _, _, _⟩, + refine ⟨⟨∑ i in s, multiset.replicate (f i / i) i, _, _⟩, _, _, _⟩, { intros i hi, simp only [exists_prop, mem_sum, mem_map, function.embedding.coe_fn_mk] at hi, rcases hi with ⟨t, ht, z⟩, apply hs, - rwa multiset.eq_of_mem_repeat z }, - { simp_rw [multiset.sum_sum, multiset.sum_repeat, nat.nsmul_eq_mul, ←hf₁], + rwa multiset.eq_of_mem_replicate z }, + { simp_rw [multiset.sum_sum, multiset.sum_replicate, nat.nsmul_eq_mul, ←hf₁], refine sum_congr rfl (λ i hi, nat.div_mul_cancel _), rcases hf₃ i hi with ⟨w, hw, hw₂⟩, rw ← hw₂, exact dvd_mul_left _ _ }, { intro i, - simp_rw [multiset.count_sum', multiset.count_repeat, sum_ite_eq], + simp_rw [multiset.count_sum', multiset.count_replicate, sum_ite_eq], split_ifs with h h, { rcases hf₃ i h with ⟨w, hw₁, hw₂⟩, rwa [← hw₂, nat.mul_div_cancel _ (hs i h)] }, @@ -355,9 +355,9 @@ begin { intros i hi, rw mem_sum at hi, rcases hi with ⟨j, hj₁, hj₂⟩, - rwa multiset.eq_of_mem_repeat hj₂ }, + rwa multiset.eq_of_mem_replicate hj₂ }, { ext i, - simp_rw [multiset.count_sum', multiset.count_repeat, sum_ite_eq], + simp_rw [multiset.count_sum', multiset.count_replicate, sum_ite_eq], split_ifs, { apply nat.div_mul_cancel, rcases hf₃ i h with ⟨w, hw, hw₂⟩, diff --git a/archive/miu_language/decision_suf.lean b/archive/miu_language/decision_suf.lean index 26b03e45d3ef8..257bea77baa2e 100644 --- a/archive/miu_language/decision_suf.lean +++ b/archive/miu_language/decision_suf.lean @@ -52,11 +52,11 @@ open miu_atom list nat We start by showing that an `miustr` `M::w` can be derived, where `w` consists only of `I`s and where `count I w` is a power of 2. -/ -private lemma der_cons_repeat (n : ℕ) : derivable (M::(repeat I (2^n))) := +private lemma der_cons_replicate (n : ℕ) : derivable (M::(replicate (2^n) I)) := begin induction n with k hk, { constructor, }, -- base case - { rw [succ_eq_add_one, pow_add, pow_one 2, mul_two,repeat_add], -- inductive step + { rw [succ_eq_add_one, pow_add, pow_one 2, mul_two,replicate_add], -- inductive step exact derivable.r2 hk, }, end @@ -81,16 +81,16 @@ an even number of `U`s and `z` is any `miustr`. Any number of successive occurrences of `"UU"` can be removed from the end of a `derivable` `miustr` to produce another `derivable` `miustr`. -/ -lemma der_of_der_append_repeat_U_even {z : miustr} {m : ℕ} (h : derivable (z ++ repeat U (m*2))) - : derivable z := +lemma der_of_der_append_replicate_U_even {z : miustr} {m : ℕ} + (h : derivable (z ++ replicate (m*2) U)) : derivable z := begin induction m with k hk, { revert h, - simp only [list.repeat, zero_mul, append_nil, imp_self], }, + simp only [list.replicate, zero_mul, append_nil, imp_self], }, { apply hk, - simp only [succ_mul, repeat_add] at h, - change repeat U 2 with [U,U] at h, - rw ←(append_nil (z ++ repeat U (k*2) )), + simp only [succ_mul, replicate_add] at h, + change replicate 2 U with [U,U] at h, + rw ←(append_nil (z ++ replicate (k*2) U)), apply derivable.r4, simp only [append_nil, append_assoc,h], }, end @@ -106,24 +106,24 @@ In fine-tuning my application of `simp`, I issued the following commend to deter We may replace several consecutive occurrences of `"III"` with the same number of `"U"`s. In application of the following lemma, `xs` will either be `[]` or `[U]`. -/ -lemma der_cons_repeat_I_repeat_U_append_of_der_cons_repeat_I_append (c k : ℕ) - (hc : c % 3 = 1 ∨ c % 3 = 2) (xs : miustr) (hder : derivable (M ::(repeat I (c+3*k)) ++ xs)) : - derivable (M::(repeat I c ++ repeat U k) ++ xs) := +lemma der_cons_replicate_I_replicate_U_append_of_der_cons_replicate_I_append (c k : ℕ) + (hc : c % 3 = 1 ∨ c % 3 = 2) (xs : miustr) (hder : derivable (M ::(replicate (c+3*k) I) ++ xs)) : + derivable (M::(replicate c I ++ replicate k U) ++ xs) := begin revert xs, induction k with a ha, - { simp only [list.repeat, mul_zero, add_zero, append_nil, forall_true_iff, imp_self],}, + { simp only [list.replicate, mul_zero, add_zero, append_nil, forall_true_iff, imp_self],}, { intro xs, specialize ha (U::xs), intro h₂, - simp only [succ_eq_add_one, repeat_add], -- We massage the goal + simp only [succ_eq_add_one, replicate_add], -- We massage the goal rw [←append_assoc, ←cons_append], -- into a form amenable - change repeat U 1 with [U], -- to the application of + change replicate 1 U with [U], -- to the application of rw [append_assoc, singleton_append], -- ha. apply ha, apply derivable.r3, - change [I,I,I] with repeat I 3, - simp only [cons_append, ←repeat_add], + change [I,I,I] with replicate 3 I, + simp only [cons_append, ←replicate_add], convert h₂, }, end @@ -178,62 +178,67 @@ end end arithmetic -lemma repeat_pow_minus_append {m : ℕ} : M :: repeat I (2^m - 1) ++ [I] = M::(repeat I (2^m)) := +lemma replicate_pow_minus_append {m : ℕ} : + M :: replicate (2^m - 1) I ++ [I] = M::(replicate (2^m) I) := begin - change [I] with repeat I 1, - rw [cons_append, ←repeat_add, tsub_add_cancel_of_le (one_le_pow' m 1)], + change [I] with replicate 1 I, + rw [cons_append, ←replicate_add, tsub_add_cancel_of_le (one_le_pow' m 1)], end /-- -`der_repeat_I_of_mod3` states that `M::y` is `derivable` if `y` is any `miustr` consisiting just of -`I`s, where `count I y` is 1 or 2 modulo 3. +`der_replicate_I_of_mod3` states that `M::y` is `derivable` if `y` is any `miustr` consisiting just +of `I`s, where `count I y` is 1 or 2 modulo 3. -/ -lemma der_repeat_I_of_mod3 (c : ℕ) (h : c % 3 = 1 ∨ c % 3 = 2): - derivable (M::(repeat I c)) := +lemma der_replicate_I_of_mod3 (c : ℕ) (h : c % 3 = 1 ∨ c % 3 = 2): + derivable (M::(replicate c I)) := begin - -- From `der_cons_repeat`, we can derive the `miustr` `M::w` described in the introduction. + -- From `der_cons_replicate`, we can derive the `miustr` `M::w` described in the introduction. cases (le_pow2_and_pow2_eq_mod3 c h) with m hm, -- `2^m` will be the number of `I`s in `M::w` - have hw₂ : derivable (M::(repeat I (2^m)) ++ repeat U ((2^m -c)/3 % 2)), + have hw₂ : derivable (M::(replicate (2^m) I) ++ replicate ((2^m -c)/3 % 2) U), { cases mod_two_eq_zero_or_one ((2^m -c)/3) with h_zero h_one, - { simp only [der_cons_repeat m, append_nil,list.repeat, h_zero], }, -- `(2^m - c)/3 ≡ 0 [MOD 2]` - { rw [h_one, ←repeat_pow_minus_append, append_assoc], -- case `(2^m - c)/3 ≡ 1 [MOD 2]` + { -- `(2^m - c)/3 ≡ 0 [MOD 2]` + simp only [der_cons_replicate m, append_nil,list.replicate, h_zero], }, + { rw [h_one, ←replicate_pow_minus_append, append_assoc], -- case `(2^m - c)/3 ≡ 1 [MOD 2]` apply derivable.r1, - rw repeat_pow_minus_append, - exact (der_cons_repeat m), }, }, - have hw₃ : derivable (M::(repeat I c) ++ repeat U ((2^m-c)/3) ++ repeat U ((2^m-c)/3 % 2)), - { apply der_cons_repeat_I_repeat_U_append_of_der_cons_repeat_I_append c ((2^m-c)/3) h, + rw replicate_pow_minus_append, + exact (der_cons_replicate m), }, }, + have hw₃ : + derivable (M::(replicate c I) ++ replicate ((2^m-c)/3) U ++ replicate ((2^m-c)/3 % 2) U), + { apply der_cons_replicate_I_replicate_U_append_of_der_cons_replicate_I_append c ((2^m-c)/3) h, convert hw₂, -- now we must show `c + 3 * ((2 ^ m - c) / 3) = 2 ^ m` rw nat.mul_div_cancel', { exact add_tsub_cancel_of_le hm.1 }, { exact (modeq_iff_dvd' hm.1).mp hm.2.symm } }, - rw [append_assoc, ←repeat_add _ _] at hw₃, + rw [append_assoc, ←replicate_add _ _] at hw₃, cases add_mod2 ((2^m-c)/3) with t ht, rw ht at hw₃, - exact der_of_der_append_repeat_U_even hw₃, + exact der_of_der_append_replicate_U_even hw₃, end example (c : ℕ) (h : c % 3 = 1 ∨ c % 3 = 2): - derivable (M::(repeat I c)) := + derivable (M::(replicate c I)) := begin - -- From `der_cons_repeat`, we can derive the `miustr` `M::w` described in the introduction. + -- From `der_cons_replicate`, we can derive the `miustr` `M::w` described in the introduction. cases (le_pow2_and_pow2_eq_mod3 c h) with m hm, -- `2^m` will be the number of `I`s in `M::w` - have hw₂ : derivable (M::(repeat I (2^m)) ++ repeat U ((2^m -c)/3 % 2)), + have hw₂ : derivable (M::(replicate (2^m) I) ++ replicate ((2^m -c)/3 % 2) U), { cases mod_two_eq_zero_or_one ((2^m -c)/3) with h_zero h_one, - { simp only [der_cons_repeat m, append_nil, list.repeat,h_zero], }, -- `(2^m - c)/3 ≡ 0 [MOD 2]` - { rw [h_one, ←repeat_pow_minus_append, append_assoc], -- case `(2^m - c)/3 ≡ 1 [MOD 2]` + { -- `(2^m - c)/3 ≡ 0 [MOD 2]` + simp only [der_cons_replicate m, append_nil, list.replicate, h_zero] }, + { rw [h_one, ←replicate_pow_minus_append, append_assoc], -- case `(2^m - c)/3 ≡ 1 [MOD 2]` apply derivable.r1, - rw repeat_pow_minus_append, - exact (der_cons_repeat m), }, }, - have hw₃ : derivable (M::(repeat I c) ++ repeat U ((2^m-c)/3) ++ repeat U ((2^m-c)/3 % 2)), - { apply der_cons_repeat_I_repeat_U_append_of_der_cons_repeat_I_append c ((2^m-c)/3) h, + rw replicate_pow_minus_append, + exact (der_cons_replicate m), }, }, + have hw₃ : + derivable (M::(replicate c I) ++ replicate ((2^m-c)/3) U ++ replicate ((2^m-c)/3 % 2) U), + { apply der_cons_replicate_I_replicate_U_append_of_der_cons_replicate_I_append c ((2^m-c)/3) h, convert hw₂, -- now we must show `c + 3 * ((2 ^ m - c) / 3) = 2 ^ m` rw nat.mul_div_cancel', { exact add_tsub_cancel_of_le hm.1 }, { exact (modeq_iff_dvd' hm.1).mp hm.2.symm } }, - rw [append_assoc, ←repeat_add _ _] at hw₃, + rw [append_assoc, ←replicate_add _ _] at hw₃, cases add_mod2 ((2^m-c)/3) with t ht, rw ht at hw₃, - exact der_of_der_append_repeat_U_even hw₃, + exact der_of_der_append_replicate_U_even hw₃, end /-! @@ -278,12 +283,12 @@ begin rcases (exists_cons_of_ne_nil this) with ⟨y,ys,rfl⟩, rw head at mhead, rw mhead at *, - rsuffices ⟨c, rfl, hc⟩ : ∃ c, repeat I c = ys ∧ (c % 3 = 1 ∨ c % 3 = 2), - { exact der_repeat_I_of_mod3 c hc, }, + rsuffices ⟨c, rfl, hc⟩ : ∃ c, replicate c I = ys ∧ (c % 3 = 1 ∨ c % 3 = 2), + { exact der_replicate_I_of_mod3 c hc, }, { simp only [count] at *, use (count I ys), refine and.intro _ hi, - apply repeat_count_eq_of_count_eq_length, + apply replicate_count_eq_of_count_eq_length, exact count_I_eq_length_of_count_U_zero_and_neg_mem hu nmtail, }, end diff --git a/src/algebra/big_operators/basic.lean b/src/algebra/big_operators/basic.lean index 07286aca97d6a..2ed573801f11f 100644 --- a/src/algebra/big_operators/basic.lean +++ b/src/algebra/big_operators/basic.lean @@ -245,7 +245,7 @@ by rw [prod_insert (not_mem_singleton.2 h), prod_singleton] @[simp, priority 1100, to_additive] lemma prod_const_one : (∏ x in s, (1 : β)) = 1 := -by simp only [finset.prod, multiset.map_const, multiset.prod_repeat, one_pow] +by simp only [finset.prod, multiset.map_const, multiset.prod_replicate, one_pow] @[simp, to_additive] lemma prod_image [decidable_eq α] {s : finset γ} {g : γ → α} : @@ -1126,7 +1126,7 @@ begin end @[simp, to_additive] lemma prod_const (b : β) : (∏ x in s, b) = b ^ s.card := -(congr_arg _ $ s.val.map_const b).trans $ multiset.prod_repeat b s.card +(congr_arg _ $ s.val.map_const b).trans $ multiset.prod_replicate b s.card @[to_additive] lemma pow_eq_prod_const (b : β) : ∀ n, b ^ n = ∏ k in range n, b := by simp @@ -1382,7 +1382,7 @@ begin classical, apply finset.induction_on' S, { simp }, intros a T haS _ haT IH, - repeat {rw finset.prod_insert haT}, + repeat { rw finset.prod_insert haT }, exact mul_dvd_mul (h a haS) IH, end diff --git a/src/algebra/big_operators/multiset/basic.lean b/src/algebra/big_operators/multiset/basic.lean index 8f44040888cba..5fa0f421bf34e 100644 --- a/src/algebra/big_operators/multiset/basic.lean +++ b/src/algebra/big_operators/multiset/basic.lean @@ -86,8 +86,8 @@ lemma prod_nsmul (m : multiset α) : ∀ (n : ℕ), (n • m).prod = m.prod ^ n | (n + 1) := by rw [add_nsmul, one_nsmul, pow_add, pow_one, prod_add, prod_nsmul n] -@[simp, to_additive] lemma prod_repeat (a : α) (n : ℕ) : (repeat a n).prod = a ^ n := -by simp [repeat, list.prod_repeat] +@[simp, to_additive] lemma prod_replicate (a : α) (n : ℕ) : (replicate n a).prod = a ^ n := +by simp [replicate, list.prod_replicate] @[to_additive] lemma prod_map_eq_pow_single [decidable_eq ι] (i : ι) (hf : ∀ i' ≠ i, i' ∈ m → f i' = 1) : @@ -107,7 +107,7 @@ end @[to_additive] lemma pow_count [decidable_eq α] (a : α) : a ^ s.count a = (s.filter (eq a)).prod := -by rw [filter_eq, prod_repeat] +by rw [filter_eq, prod_replicate] @[to_additive] lemma prod_hom [comm_monoid β] (s : multiset α) {F : Type*} [monoid_hom_class F α β] (f : F) : @@ -134,7 +134,7 @@ quotient.induction_on s $ λ l, by simp only [l.prod_hom_rel h₁ h₂, quot_mk_to_coe, coe_map, coe_prod] @[to_additive] -lemma prod_map_one : prod (m.map (λ i, (1 : α))) = 1 := by rw [map_const, prod_repeat, one_pow] +lemma prod_map_one : prod (m.map (λ i, (1 : α))) = 1 := by rw [map_const, prod_replicate, one_pow] @[simp, to_additive] lemma prod_map_mul : (m.map $ λ i, f i * g i).prod = (m.map f).prod * (m.map g).prod := @@ -321,7 +321,7 @@ lemma prod_le_prod_map (f : α → α) (h : ∀ x, x ∈ s → x ≤ f x) : s.pr @[to_additive card_nsmul_le_sum] lemma pow_card_le_prod (h : ∀ x ∈ s, a ≤ x) : a ^ s.card ≤ s.prod := -by { rw [←multiset.prod_repeat, ←multiset.map_const], exact prod_map_le_prod _ h } +by { rw [←multiset.prod_replicate, ←multiset.map_const], exact prod_map_le_prod _ h } end ordered_comm_monoid diff --git a/src/algebra/gcd_monoid/multiset.lean b/src/algebra/gcd_monoid/multiset.lean index a5979fe6249a2..1647017c603c6 100644 --- a/src/algebra/gcd_monoid/multiset.lean +++ b/src/algebra/gcd_monoid/multiset.lean @@ -180,8 +180,8 @@ lemma extract_gcd (s : multiset α) (hs : s ≠ 0) : begin classical, by_cases h : ∀ x ∈ s, x = (0 : α), - { use repeat 1 s.card, - rw [map_repeat, eq_repeat, mul_one, s.gcd_eq_zero_iff.2 h, ←nsmul_singleton, ←gcd_dedup], + { use replicate s.card 1, + rw [map_replicate, eq_replicate, mul_one, s.gcd_eq_zero_iff.2 h, ←nsmul_singleton, ←gcd_dedup], rw [dedup_nsmul (card_pos.2 hs).ne', dedup_singleton, gcd_singleton], exact ⟨⟨rfl, h⟩, normalize_one⟩ }, { choose f hf using @gcd_dvd _ _ _ s, diff --git a/src/algebra/hom/freiman.lean b/src/algebra/hom/freiman.lean index f55b01bfd49b5..3c32ac4ab79a4 100644 --- a/src/algebra/hom/freiman.lean +++ b/src/algebra/hom/freiman.lean @@ -196,7 +196,7 @@ ext $ λ x, rfl def const (A : set α) (n : ℕ) (b : β) : A →*[n] β := { to_fun := λ _, b, map_prod_eq_map_prod' := λ s t _ _ hs ht _, - by rw [multiset.map_const, multiset.map_const, prod_repeat, prod_repeat, hs, ht] } + by rw [multiset.map_const, multiset.map_const, prod_replicate, prod_replicate, hs, ht] } @[simp, to_additive] lemma const_apply (n : ℕ) (b : β) (x : α) : const A n b x = b := rfl @@ -341,7 +341,7 @@ begin rw [hs, ht] }, rw [←hs, card_pos_iff_exists_mem] at hm, obtain ⟨a, ha⟩ := hm, - suffices : ((s + repeat a (n - m)).map f).prod = ((t + repeat a (n - m)).map f).prod, + suffices : ((s + replicate (n - m) a).map f).prod = ((t + replicate (n - m) a).map f).prod, { simp_rw [multiset.map_add, prod_add] at this, exact mul_right_cancel this }, replace ha := hsA _ ha, @@ -349,12 +349,12 @@ begin rotate 2, assumption, -- Can't infer `A` and `n` from the context, so do it manually. { rw mem_add at hx, refine hx.elim (hsA _) (λ h, _), - rwa eq_of_mem_repeat h }, + rwa eq_of_mem_replicate h }, { rw mem_add at hx, refine hx.elim (htA _) (λ h, _), - rwa eq_of_mem_repeat h }, - { rw [card_add, hs, card_repeat, add_tsub_cancel_of_le h] }, - { rw [card_add, ht, card_repeat, add_tsub_cancel_of_le h] }, + rwa eq_of_mem_replicate h }, + { rw [card_add, hs, card_replicate, add_tsub_cancel_of_le h] }, + { rw [card_add, ht, card_replicate, add_tsub_cancel_of_le h] }, { rw [prod_add, prod_add, hst] } end diff --git a/src/algebra/module/dedekind_domain.lean b/src/algebra/module/dedekind_domain.lean index 775f0ed143055..db6de92fae59f 100644 --- a/src/algebra/module/dedekind_domain.lean +++ b/src/algebra/module/dedekind_domain.lean @@ -48,7 +48,7 @@ begin { suffices : (normalized_factors _).count p = 0, { rw [this, zero_min, pow_zero, ideal.one_eq_top] }, { rw [multiset.count_eq_zero, normalized_factors_of_irreducible_pow - (prime_of_mem q hq).irreducible, multiset.mem_repeat], + (prime_of_mem q hq).irreducible, multiset.mem_replicate], exact λ H, pq $ H.2.trans $ normalize_eq q } }, { rw ← ideal.zero_eq_bot, apply pow_ne_zero, exact (prime_of_mem q hq).ne_zero }, { exact (prime_of_mem p hp).irreducible } } diff --git a/src/algebra/monoid_algebra/degree.lean b/src/algebra/monoid_algebra/degree.lean index 831bc4ae4777e..602c3a8fa079f 100644 --- a/src/algebra/monoid_algebra/degree.lean +++ b/src/algebra/monoid_algebra/degree.lean @@ -96,9 +96,9 @@ lemma sup_support_pow_le (degb0 : degb 0 ≤ 0) (degbm : ∀ a b, degb (a + b) (n : ℕ) (f : add_monoid_algebra R A) : (f ^ n).support.sup degb ≤ n • (f.support.sup degb) := begin - rw [← list.prod_repeat, ←list.sum_repeat], + rw [← list.prod_replicate, ←list.sum_replicate], refine (sup_support_list_prod_le degb0 degbm _).trans_eq _, - rw list.map_repeat, + rw list.map_replicate, end lemma le_inf_support_pow (degt0 : 0 ≤ degt 0) (degtm : ∀ a b, degt a + degt b ≤ degt (a + b)) diff --git a/src/algebra/polynomial/big_operators.lean b/src/algebra/polynomial/big_operators.lean index 87967ccbfdc9d..35f66c3e060ed 100644 --- a/src/algebra/polynomial/big_operators.lean +++ b/src/algebra/polynomial/big_operators.lean @@ -192,8 +192,8 @@ begin nontriviality R, apply nat_degree_multiset_prod', suffices : (t.map (λ f, leading_coeff f)).prod = 1, { rw this, simp }, - convert prod_repeat (1 : R) t.card, - { simp only [eq_repeat, multiset.card_map, eq_self_iff_true, true_and], + convert prod_replicate (1 : R) t.card, + { simp only [eq_replicate, multiset.card_map, eq_self_iff_true, true_and], rintros i hi, obtain ⟨i, hi, rfl⟩ := multiset.mem_map.mp hi, apply h, assumption }, diff --git a/src/combinatorics/composition.lean b/src/combinatorics/composition.lean index b39773ebcc5cb..c7fad7f0386ec 100644 --- a/src/combinatorics/composition.lean +++ b/src/combinatorics/composition.lean @@ -427,22 +427,22 @@ end /-- The composition made of blocks all of size `1`. -/ def ones (n : ℕ) : composition n := -⟨repeat (1 : ℕ) n, λ i hi, by simp [list.eq_of_mem_repeat hi], by simp⟩ +⟨replicate n (1 : ℕ), λ i hi, by simp [list.eq_of_mem_replicate hi], by simp⟩ instance {n : ℕ} : inhabited (composition n) := ⟨composition.ones n⟩ @[simp] lemma ones_length (n : ℕ) : (ones n).length = n := -list.length_repeat 1 n +list.length_replicate n 1 -@[simp] lemma ones_blocks (n : ℕ) : (ones n).blocks = repeat (1 : ℕ) n := rfl +@[simp] lemma ones_blocks (n : ℕ) : (ones n).blocks = replicate n (1 : ℕ) := rfl @[simp] lemma ones_blocks_fun (n : ℕ) (i : fin (ones n).length) : (ones n).blocks_fun i = 1 := by simp [blocks_fun, ones, blocks, i.2] @[simp] lemma ones_size_up_to (n : ℕ) (i : ℕ) : (ones n).size_up_to i = min i n := -by simp [size_up_to, ones_blocks, take_repeat] +by simp [size_up_to, ones_blocks, take_replicate] @[simp] lemma ones_embedding (i : fin (ones n).length) (h : 0 < (ones n).blocks_fun i) : (ones n).embedding i ⟨0, h⟩ = ⟨i, lt_of_lt_of_le i.2 (ones n).length_le⟩ := @@ -453,10 +453,10 @@ lemma eq_ones_iff {c : composition n} : begin split, { rintro rfl, - exact λ i, eq_of_mem_repeat }, + exact λ i, eq_of_mem_replicate }, { assume H, ext1, - have A : c.blocks = repeat 1 c.blocks.length := eq_repeat_of_mem H, + have A : c.blocks = replicate c.blocks.length 1 := eq_replicate_of_mem H, have : c.blocks.length = n, by { conv_rhs { rw [← c.blocks_sum, A] }, simp }, rw [A, this, ones_blocks] }, end diff --git a/src/computability/turing_machine.lean b/src/computability/turing_machine.lean index 6fe07d75f83b5..7924f8ed6a3a4 100644 --- a/src/computability/turing_machine.lean +++ b/src/computability/turing_machine.lean @@ -67,22 +67,22 @@ namespace turing /-- The `blank_extends` partial order holds of `l₁` and `l₂` if `l₂` is obtained by adding blanks (`default : Γ`) to the end of `l₁`. -/ def blank_extends {Γ} [inhabited Γ] (l₁ l₂ : list Γ) : Prop := -∃ n, l₂ = l₁ ++ list.repeat default n +∃ n, l₂ = l₁ ++ list.replicate n default @[refl] theorem blank_extends.refl {Γ} [inhabited Γ] (l : list Γ) : blank_extends l l := ⟨0, by simp⟩ @[trans] theorem blank_extends.trans {Γ} [inhabited Γ] {l₁ l₂ l₃ : list Γ} : blank_extends l₁ l₂ → blank_extends l₂ l₃ → blank_extends l₁ l₃ := -by { rintro ⟨i, rfl⟩ ⟨j, rfl⟩, exact ⟨i+j, by simp [list.repeat_add]⟩ } +by { rintro ⟨i, rfl⟩ ⟨j, rfl⟩, exact ⟨i+j, by simp [list.replicate_add]⟩ } theorem blank_extends.below_of_le {Γ} [inhabited Γ] {l l₁ l₂ : list Γ} : blank_extends l l₁ → blank_extends l l₂ → l₁.length ≤ l₂.length → blank_extends l₁ l₂ := begin rintro ⟨i, rfl⟩ ⟨j, rfl⟩ h, use j - i, - simp only [list.length_append, add_le_add_iff_left, list.length_repeat] at h, - simp only [← list.repeat_add, add_tsub_cancel_of_le h, list.append_assoc], + simp only [list.length_append, add_le_add_iff_left, list.length_replicate] at h, + simp only [← list.replicate_add, add_tsub_cancel_of_le h, list.append_assoc], end /-- Any two extensions by blank `l₁,l₂` of `l` have a common join (which can be taken to be the @@ -101,9 +101,9 @@ theorem blank_extends.above_of_le {Γ} [inhabited Γ] {l l₁ l₂ : list Γ} : begin rintro ⟨i, rfl⟩ ⟨j, e⟩ h, use i - j, refine list.append_right_cancel (e.symm.trans _), - rw [list.append_assoc, ← list.repeat_add, tsub_add_cancel_of_le], + rw [list.append_assoc, ← list.replicate_add, tsub_add_cancel_of_le], apply_fun list.length at e, - simp only [list.length_append, list.length_repeat] at e, + simp only [list.length_append, list.length_replicate] at e, rwa [← add_le_add_iff_left, e, add_le_add_iff_right] end @@ -245,7 +245,7 @@ l.lift_on (λ l, list.inth l n) begin cases lt_or_le _ _ with h h, {rw list.inth_append _ _ _ h}, rw list.inth_eq_default _ h, cases le_or_lt _ _ with h₂ h₂, {rw list.inth_eq_default _ h₂}, - rw [list.inth_eq_nth_le _ h₂, list.nth_le_append_right h, list.nth_le_repeat] + rw [list.inth_eq_nth_le _ h₂, list.nth_le_append_right h, list.nth_le_replicate] end @[simp] theorem list_blank.nth_mk {Γ} [inhabited Γ] (l : list Γ) (n : ℕ) : @@ -272,12 +272,12 @@ begin swap, { exact (this $ λ i, (H i).symm).symm }, refine quotient.sound' (or.inl ⟨l₂.length - l₁.length, _⟩), refine list.ext_le _ (λ i h h₂, eq.symm _), - { simp only [add_tsub_cancel_of_le h, list.length_append, list.length_repeat] }, + { simp only [add_tsub_cancel_of_le h, list.length_append, list.length_replicate] }, simp only [list_blank.nth_mk] at H, cases lt_or_le i l₁.length with h' h', { simp only [list.nth_le_append _ h', list.nth_le_nth h, list.nth_le_nth h', ←list.inth_eq_nth_le _ h, ←list.inth_eq_nth_le _ h', H] }, - { simp only [list.nth_le_append_right h', list.nth_le_repeat, list.nth_le_nth h, + { simp only [list.nth_le_append_right h', list.nth_le_replicate, list.nth_le_nth h, list.nth_len_le h', ←list.inth_eq_default _ h', H, list.inth_eq_nth_le _ h] } end @@ -326,7 +326,7 @@ def list_blank.map {Γ Γ'} [inhabited Γ] [inhabited Γ'] (f : pointed_map Γ Γ') (l : list_blank Γ) : list_blank Γ' := l.lift_on (λ l, list_blank.mk (list.map f l)) begin rintro l _ ⟨i, rfl⟩, refine quotient.sound' (or.inl ⟨i, _⟩), - simp only [pointed_map.map_pt, list.map_append, list.map_repeat], + simp only [pointed_map.map_pt, list.map_append, list.map_replicate], end @[simp] theorem list_blank.map_mk {Γ Γ'} [inhabited Γ] [inhabited Γ'] @@ -393,12 +393,12 @@ l₃.induction_on $ by intro; simp only [list_blank.append_mk, list.append_assoc is sent to a sequence of default elements. -/ def list_blank.bind {Γ Γ'} [inhabited Γ] [inhabited Γ'] (l : list_blank Γ) (f : Γ → list Γ') - (hf : ∃ n, f default = list.repeat default n) : list_blank Γ' := + (hf : ∃ n, f default = list.replicate n default) : list_blank Γ' := l.lift_on (λ l, list_blank.mk (list.bind l f)) begin rintro l _ ⟨i, rfl⟩, cases hf with n e, refine quotient.sound' (or.inl ⟨i * n, _⟩), rw [list.bind_append, mul_comm], congr, induction i with i IH, refl, - simp only [IH, e, list.repeat_add, nat.mul_succ, add_comm, list.repeat_succ, list.cons_bind], + simp only [IH, e, list.replicate_add, nat.mul_succ, add_comm, list.replicate_succ, list.cons_bind] end @[simp] lemma list_blank.bind_mk {Γ Γ'} [inhabited Γ] [inhabited Γ'] @@ -1365,7 +1365,7 @@ parameters {Γ : Type*} [inhabited Γ] theorem exists_enc_dec [fintype Γ] : ∃ n (enc : Γ → vector bool n) (dec : vector bool n → Γ), - enc default = vector.repeat ff n ∧ ∀ a, dec (enc a) = a := + enc default = vector.replicate n ff ∧ ∀ a, dec (enc a) = a := begin letI := classical.dec_eq Γ, let n := fintype.card Γ, @@ -1375,7 +1375,7 @@ begin let H := (F.to_embedding.trans G).trans (equiv.vector_equiv_fin _ _).symm.to_embedding, classical, - let enc := H.set_value default (vector.repeat ff n), + let enc := H.set_value default (vector.replicate n ff), exact ⟨_, enc, function.inv_fun enc, H.set_value_eq _ _, function.left_inverse_inv_fun enc.2⟩ end @@ -1457,7 +1457,7 @@ from λ f hf, this n _ (by intro; simp only [supports_stmt_move, hf]), split; apply IH; intro; apply hf, end -parameter (enc0 : enc default = vector.repeat ff n) +parameter (enc0 : enc default = vector.replicate n ff) section parameter {enc} @@ -1469,8 +1469,8 @@ begin refine tape.mk' (L.bind (λ x, (enc x).to_list.reverse) ⟨n, _⟩) (R.bind (λ x, (enc x).to_list) ⟨n, _⟩); - simp only [enc0, vector.repeat, - list.reverse_repeat, bool.default_bool, vector.to_list_mk] + simp only [enc0, vector.replicate, + list.reverse_replicate, bool.default_bool, vector.to_list_mk] end /-- The low level tape corresponding to the given tape over alphabet `Γ`. -/ diff --git a/src/data/bitvec/core.lean b/src/data/bitvec/core.lean index 7f0c665168f76..8b3fbda636778 100644 --- a/src/data/bitvec/core.lean +++ b/src/data/bitvec/core.lean @@ -26,12 +26,12 @@ open vector local infix `++ₜ`:65 := vector.append /-- Create a zero bitvector -/ -@[reducible] protected def zero (n : ℕ) : bitvec n := repeat ff n +@[reducible] protected def zero (n : ℕ) : bitvec n := replicate n ff /-- Create a bitvector of length `n` whose `n-1`st entry is 1 and other entries are 0. -/ @[reducible] protected def one : Π (n : ℕ), bitvec n | 0 := nil -| (succ n) := repeat ff n ++ₜ tt::ᵥnil +| (succ n) := replicate n ff ++ₜ tt::ᵥnil /-- Create a bitvector from another with a provably equal length. -/ protected def cong {a b : ℕ} (h : a = b) : bitvec a → bitvec b @@ -49,7 +49,7 @@ variable {n : ℕ} If `x.length < i` then this will return the all-`ff`s bitvector. -/ def shl (x : bitvec n) (i : ℕ) : bitvec n := bitvec.cong (by simp) $ - drop i x ++ₜ repeat ff (min n i) + drop i x ++ₜ replicate (min n i) ff /-- `fill_shr x i fill` is the bitvector obtained by right-shifting `x` `i` times and then padding with `fill : bool`. If `x.length < i` then this will return the constant `fill` @@ -64,7 +64,7 @@ bitvec.cong { have h₁ := le_of_not_ge h, rw [min_eq_left h₁, tsub_eq_zero_iff_le.mpr h₁, zero_min, nat.add_zero] } end $ - repeat fill (min n i) ++ₜ take (n-i) x + replicate (min n i) fill ++ₜ take (n-i) x /-- unsigned shift right -/ def ushr (x : bitvec n) (i : ℕ) : bitvec n := diff --git a/src/data/dfinsupp/multiset.lean b/src/data/dfinsupp/multiset.lean index 3d1da56331209..ade908fd4639b 100644 --- a/src/data/dfinsupp/multiset.lean +++ b/src/data/dfinsupp/multiset.lean @@ -26,10 +26,10 @@ variables [decidable_eq α] /-- A computable version of `finsupp.to_multiset`. -/ def to_multiset : (Π₀ a : α, ℕ) →+ multiset α := -dfinsupp.sum_add_hom (λ a : α, multiset.repeat_add_monoid_hom a) +dfinsupp.sum_add_hom (λ a : α, multiset.replicate_add_monoid_hom a) @[simp] lemma to_multiset_single (a : α) (n : ℕ) : - to_multiset (dfinsupp.single a n) = multiset.repeat a n := + to_multiset (dfinsupp.single a n) = multiset.replicate n a := dfinsupp.sum_add_hom_single _ _ _ end dfinsupp @@ -52,16 +52,16 @@ def to_dfinsupp : multiset α →+ Π₀ a : α, ℕ := s.to_dfinsupp.support = s.to_finset := (finset.filter_eq_self _).mpr (λ x hx, count_ne_zero.mpr $ multiset.mem_to_finset.1 hx) -@[simp] lemma to_dfinsupp_repeat (a : α) (n : ℕ) : - to_dfinsupp (multiset.repeat a n) = dfinsupp.single a n := +@[simp] lemma to_dfinsupp_replicate (a : α) (n : ℕ) : + to_dfinsupp (multiset.replicate n a) = dfinsupp.single a n := begin ext i, dsimp [to_dfinsupp], - simp [count_repeat, eq_comm], + simp [count_replicate, eq_comm], end @[simp] lemma to_dfinsupp_singleton (a : α) : to_dfinsupp {a} = dfinsupp.single a 1 := -by rw [←repeat_one, to_dfinsupp_repeat] +by rw [←replicate_one, to_dfinsupp_replicate] /-- `multiset.to_dfinsupp` as an `add_equiv`. -/ @[simps apply symm_apply] diff --git a/src/data/finset/basic.lean b/src/data/finset/basic.lean index 0966010e09b20..af99e841e5926 100644 --- a/src/data/finset/basic.lean +++ b/src/data/finset/basic.lean @@ -2163,8 +2163,8 @@ end @[simp] lemma to_finset_reverse {l : list α} : to_finset l.reverse = l.to_finset := to_finset_eq_of_perm _ _ (reverse_perm l) -lemma to_finset_repeat_of_ne_zero {n : ℕ} (hn : n ≠ 0) : (list.repeat a n).to_finset = {a} := -by { ext x, simp [hn, list.mem_repeat] } +lemma to_finset_replicate_of_ne_zero {n : ℕ} (hn : n ≠ 0) : (list.replicate n a).to_finset = {a} := +by { ext x, simp [hn, list.mem_replicate] } @[simp] lemma to_finset_union (l l' : list α) : (l ∪ l').to_finset = l.to_finset ∪ l'.to_finset := by { ext, simp } diff --git a/src/data/finset/sym.lean b/src/data/finset/sym.lean index b023650775f22..408abf97802cf 100644 --- a/src/data/finset/sym.lean +++ b/src/data/finset/sym.lean @@ -113,15 +113,15 @@ end @[simp] lemma sym_empty (n : ℕ) : (∅ : finset α).sym (n + 1) = ∅ := rfl -lemma repeat_mem_sym (ha : a ∈ s) (n : ℕ) : sym.repeat a n ∈ s.sym n := -mem_sym_iff.2 $ λ b hb, by rwa (sym.mem_repeat.1 hb).2 +lemma replicate_mem_sym (ha : a ∈ s) (n : ℕ) : sym.replicate n a ∈ s.sym n := +mem_sym_iff.2 $ λ b hb, by rwa (sym.mem_replicate.1 hb).2 protected lemma nonempty.sym (h : s.nonempty) (n : ℕ) : (s.sym n).nonempty := -let ⟨a, ha⟩ := h in ⟨_, repeat_mem_sym ha n⟩ +let ⟨a, ha⟩ := h in ⟨_, replicate_mem_sym ha n⟩ -@[simp] lemma sym_singleton (a : α) (n : ℕ) : ({a} : finset α).sym n = {sym.repeat a n} := -eq_singleton_iff_nonempty_unique_mem.2 ⟨(singleton_nonempty _).sym n, - λ s hs, sym.eq_repeat_iff.2 $ λ b hb, eq_of_mem_singleton $ mem_sym_iff.1 hs _ hb⟩ +@[simp] lemma sym_singleton (a : α) (n : ℕ) : ({a} : finset α).sym n = {sym.replicate n a} := +eq_singleton_iff_unique_mem.2 ⟨replicate_mem_sym (mem_singleton.2 rfl) _, + λ s hs, sym.eq_replicate_iff.2 $ λ b hb, eq_of_mem_singleton $ mem_sym_iff.1 hs _ hb⟩ lemma eq_empty_of_sym_eq_empty (h : s.sym n = ∅) : s = ∅ := begin diff --git a/src/data/fintype/card.lean b/src/data/fintype/card.lean index 5f8f084613e0e..9deaf0aa704ef 100644 --- a/src/data/fintype/card.lean +++ b/src/data/fintype/card.lean @@ -822,7 +822,8 @@ instance : infinite ℤ := infinite.of_injective int.of_nat (λ _ _, int.of_nat.inj) instance [nonempty α] : infinite (multiset α) := -let ⟨x⟩ := ‹nonempty α› in infinite.of_injective (multiset.repeat x) (multiset.repeat_injective _) +let ⟨x⟩ := ‹nonempty α› in + infinite.of_injective (λ n, multiset.replicate n x) (multiset.replicate_left_injective _) instance [nonempty α] : infinite (list α) := infinite.of_surjective (coe : list α → multiset α) (surjective_quot_mk _) diff --git a/src/data/list/basic.lean b/src/data/list/basic.lean index 913dbfa253faf..93f0a0ec958a6 100644 --- a/src/data/list/basic.lean +++ b/src/data/list/basic.lean @@ -473,72 +473,68 @@ begin apply nat.le_add_right end -/-! ### repeat -/ +/-! ### replicate -/ -@[simp] theorem repeat_succ (a : α) (n) : repeat a (n + 1) = a :: repeat a n := rfl +@[simp] theorem replicate_succ (a : α) (n) : replicate (n + 1) a = a :: replicate n a := rfl -theorem mem_repeat {a b : α} : ∀ {n}, b ∈ repeat a n ↔ n ≠ 0 ∧ b = a +@[simp] theorem length_replicate : ∀ n (a : α), length (replicate n a) = n +| 0 a := rfl +| (n + 1) a := congr_arg nat.succ (length_replicate n a) + +theorem mem_replicate {a b : α} : ∀ {n}, b ∈ replicate n a ↔ n ≠ 0 ∧ b = a | 0 := by simp -| (n + 1) := by simp [mem_repeat] +| (n + 1) := by simp [mem_replicate] -theorem eq_of_mem_repeat {a b : α} {n} (h : b ∈ repeat a n) : b = a := -(mem_repeat.1 h).2 +theorem eq_of_mem_replicate {a b : α} {n} (h : b ∈ replicate n a) : b = a := +(mem_replicate.1 h).2 -theorem eq_repeat_of_mem {a : α} : ∀ {l : list α}, (∀ b ∈ l, b = a) → l = repeat a l.length -| [] H := rfl -| (b::l) H := by cases forall_mem_cons.1 H with H₁ H₂; - unfold length repeat; congr; [exact H₁, exact eq_repeat_of_mem H₂] +theorem eq_replicate_length {a : α} : ∀ {l : list α}, l = replicate l.length a ↔ ∀ b ∈ l, b = a +| [] := by simp +| (b :: l) := by simp [eq_replicate_length] -theorem eq_repeat' {a : α} {l : list α} : l = repeat a l.length ↔ ∀ b ∈ l, b = a := -⟨λ h, h.symm ▸ λ b, eq_of_mem_repeat, eq_repeat_of_mem⟩ +alias eq_replicate_length ↔ _ eq_replicate_of_mem -theorem eq_repeat {a : α} {n} {l : list α} : l = repeat a n ↔ length l = n ∧ ∀ b ∈ l, b = a := -⟨λ h, h.symm ▸ ⟨length_repeat _ _, λ b, eq_of_mem_repeat⟩, - λ ⟨e, al⟩, e ▸ eq_repeat_of_mem al⟩ +theorem eq_replicate {a : α} {n} {l : list α} : l = replicate n a ↔ length l = n ∧ ∀ b ∈ l, b = a := +⟨λ h, h.symm ▸ ⟨length_replicate _ _, λ b, eq_of_mem_replicate⟩, + λ ⟨e, al⟩, e ▸ eq_replicate_of_mem al⟩ -theorem repeat_add (a : α) (m n) : repeat a (m + n) = repeat a m ++ repeat a n := -by induction m; simp only [*, zero_add, succ_add, repeat]; split; refl +theorem replicate_add (a : α) (m n) : replicate (m + n) a = replicate m a ++ replicate n a := +by induction m; simp only [*, zero_add, succ_add, replicate]; refl -theorem repeat_subset_singleton (a : α) (n) : repeat a n ⊆ [a] := -λ b h, mem_singleton.2 (eq_of_mem_repeat h) +theorem replicate_subset_singleton (a : α) (n) : replicate n a ⊆ [a] := +λ b h, mem_singleton.2 (eq_of_mem_replicate h) -lemma subset_singleton_iff {a : α} : ∀ L : list α, L ⊆ [a] ↔ ∃ n, L = repeat a n -| [] := ⟨λ h, ⟨0, by simp⟩, by simp⟩ -| (h :: L) := begin - refine ⟨λ h, _, λ ⟨k, hk⟩, by simp [hk, repeat_subset_singleton]⟩, - rw [cons_subset] at h, - obtain ⟨n, rfl⟩ := (subset_singleton_iff L).mp h.2, - exact ⟨n.succ, by simp [mem_singleton.mp h.1]⟩ -end +lemma subset_singleton_iff {a : α} {L : list α} : L ⊆ [a] ↔ ∃ n, L = replicate n a := +by simp only [eq_replicate, subset_def, mem_singleton, exists_eq_left'] -@[simp] theorem map_repeat (f : α → β) (a : α) (n) : map f (repeat a n) = repeat (f a) n := -by induction n; [refl, simp only [*, repeat, map]]; split; refl +@[simp] theorem map_replicate (f : α → β) (a : α) (n) : map f (replicate n a) = replicate n (f a) := +by induction n; [refl, simp only [*, replicate, map]]; split; refl -@[simp] theorem tail_repeat (a : α) (n) : tail (repeat a n) = repeat a n.pred := +@[simp] theorem tail_replicate (n) (a : α) : tail (replicate n a) = replicate (n - 1) a := by cases n; refl -@[simp] theorem join_repeat_nil (n : ℕ) : join (repeat [] n) = @nil α := -by induction n; [refl, simp only [*, repeat, join, append_nil]] +@[simp] theorem join_replicate_nil (n : ℕ) : join (replicate n []) = @nil α := +by induction n; [refl, simp only [*, replicate, join, append_nil]] -lemma repeat_left_injective {n : ℕ} (hn : n ≠ 0) : - function.injective (λ a : α, repeat a n) := -λ a b h, (eq_repeat.1 h).2 _ $ mem_repeat.2 ⟨hn, rfl⟩ +lemma replicate_right_injective {n : ℕ} (hn : n ≠ 0) : + function.injective (replicate n : α → list α) := +λ _ _ h, (eq_replicate.1 h).2 _ $ mem_replicate.2 ⟨hn, rfl⟩ -lemma repeat_left_inj {a b : α} {n : ℕ} (hn : n ≠ 0) : - repeat a n = repeat b n ↔ a = b := -(repeat_left_injective hn).eq_iff +lemma replicate_right_inj {a b : α} {n : ℕ} (hn : n ≠ 0) : + replicate n a = replicate n b ↔ a = b := +(replicate_right_injective hn).eq_iff -@[simp] lemma repeat_left_inj' {a b : α} : - ∀ {n}, repeat a n = repeat b n ↔ n = 0 ∨ a = b +@[simp] lemma replicate_right_inj' {a b : α} : + ∀ {n}, replicate n a = replicate n b ↔ n = 0 ∨ a = b | 0 := by simp -| (n + 1) := (repeat_left_inj n.succ_ne_zero).trans $ by simp only [n.succ_ne_zero, false_or] +| (n + 1) := (replicate_right_inj n.succ_ne_zero).trans $ by simp only [n.succ_ne_zero, false_or] -lemma repeat_right_injective (a : α) : function.injective (repeat a) := -function.left_inverse.injective (length_repeat a) +lemma replicate_left_injective (a : α) : function.injective (λ n, replicate n a) := +function.left_inverse.injective (λ n, length_replicate n a) -@[simp] lemma repeat_right_inj {a : α} {n m : ℕ} : - repeat a n = repeat a m ↔ n = m := -(repeat_right_injective a).eq_iff +@[simp] lemma replicate_left_inj {a : α} {n m : ℕ} : + replicate n a = replicate m a ↔ n = m := +(replicate_left_injective a).eq_iff /-! ### pure -/ @@ -664,9 +660,9 @@ by simp only [reverse_core_eq, map_append, map_reverse] by induction l; [refl, simp only [*, reverse_cons, mem_append, mem_singleton, mem_cons_iff, not_mem_nil, false_or, or_false, or_comm]] -@[simp] theorem reverse_repeat (a : α) (n) : reverse (repeat a n) = repeat a n := -eq_repeat.2 ⟨by simp only [length_reverse, length_repeat], - λ b h, eq_of_mem_repeat (mem_reverse.1 h)⟩ +@[simp] theorem reverse_replicate (a : α) (n) : reverse (replicate n a) = replicate n a := +eq_replicate.2 ⟨by simp only [length_reverse, length_replicate], + λ b h, eq_of_mem_replicate (mem_reverse.1 h)⟩ /-! ### empty -/ @@ -733,13 +729,12 @@ theorem last_mem : ∀ {l : list α} (h : l ≠ []), last l h ∈ l | [a] h := or.inl rfl | (a::b::l) h := or.inr $ by { rw [last_cons_cons], exact last_mem (cons_ne_nil b l) } -lemma last_repeat_succ (a m : ℕ) : - (repeat a m.succ).last (ne_nil_of_length_eq_succ - (show (repeat a m.succ).length = m.succ, by rw length_repeat)) = a := +lemma last_replicate_succ (a m : ℕ) : + (replicate (m + 1) a).last (ne_nil_of_length_eq_succ (length_replicate (m + 1) a)) = a := begin induction m with k IH, { simp }, - { simpa only [repeat_succ, last] } + { simpa only [replicate_succ, last] } end /-! ### last' -/ @@ -1068,15 +1063,16 @@ eq_nil_of_subset_nil $ s.subset @[simp] theorem sublist_nil_iff_eq_nil {l : list α} : l <+ [] ↔ l = [] := ⟨eq_nil_of_sublist_nil, λ H, H ▸ sublist.refl _⟩ -@[simp] theorem repeat_sublist_repeat (a : α) {m n} : repeat a m <+ repeat a n ↔ m ≤ n := -⟨λ h, by simpa only [length_repeat] using h.length_le, - λ h, by induction h; [refl, simp only [*, repeat_succ, sublist.cons]] ⟩ +@[simp] theorem replicate_sublist_replicate (a : α) {m n} : + replicate m a <+ replicate n a ↔ m ≤ n := +⟨λ h, by simpa only [length_replicate] using h.length_le, + λ h, by induction h; [refl, simp only [*, replicate_succ, sublist.cons]] ⟩ -lemma sublist_repeat_iff {l : list α} {a : α} {n : ℕ} : - l <+ repeat a n ↔ ∃ (k ≤ n), l = repeat a k := -⟨λ h, ⟨l.length, h.length_le.trans (length_repeat _ _).le, eq_repeat.mpr - ⟨rfl, λ b hb, list.eq_of_mem_repeat (h.subset hb)⟩⟩, - by { rintro ⟨k, h, rfl⟩, exact (repeat_sublist_repeat _).mpr h }⟩ +lemma sublist_replicate_iff {l : list α} {a : α} {n : ℕ} : + l <+ replicate n a ↔ ∃ k ≤ n, l = replicate k a := +⟨λ h, ⟨l.length, h.length_le.trans (length_replicate _ _).le, eq_replicate_length.mpr $ + λ b hb, eq_of_mem_replicate (h.subset hb)⟩, + by { rintro ⟨k, h, rfl⟩, exact (replicate_sublist_replicate _).mpr h }⟩ theorem sublist.eq_of_length : ∀ {l₁ l₂ : list α}, l₁ <+ l₂ → length l₁ = length l₂ → l₁ = l₂ | ._ ._ sublist.slnil h := rfl @@ -1303,9 +1299,9 @@ lemma nth_le_append_right : ∀ {l₁ l₂ : list α} {n : ℕ} (h₁ : l₁.len rw nth_le_append_right (nat.lt_succ_iff.mp h₁), end -@[simp] lemma nth_le_repeat (a : α) {n m : ℕ} (h : m < (list.repeat a n).length) : - (list.repeat a n).nth_le m h = a := -eq_of_mem_repeat (nth_le_mem _ _ _) +@[simp] lemma nth_le_replicate (a : α) {n m : ℕ} (h : m < (list.replicate n a).length) : + (list.replicate n a).nth_le m h = a := +eq_of_mem_replicate (nth_le_mem _ _ _) lemma nth_append {l₁ l₂ : list α} {n : ℕ} (hn : n < l₁.length) : (l₁ ++ l₂).nth n = l₁.nth n := @@ -1796,22 +1792,17 @@ begin { simpa using l_ih } } end -lemma map_eq_repeat_iff {l : list α} {f : α → β} {b : β} : - l.map f = repeat b l.length ↔ (∀ x ∈ l, f x = b) := -begin - induction l with x l' ih, - { simp only [repeat, length, not_mem_nil, is_empty.forall_iff, implies_true_iff, - map_nil, eq_self_iff_true], }, - { simp only [map, length, mem_cons_iff, forall_eq_or_imp, repeat_succ, and.congr_right_iff], - exact λ _, ih, } -end +lemma map_eq_replicate_iff {l : list α} {f : α → β} {b : β} : + l.map f = replicate l.length b ↔ (∀ x ∈ l, f x = b) := +by simp [eq_replicate] -@[simp] theorem map_const (l : list α) (b : β) : map (function.const α b) l = repeat b l.length := -map_eq_repeat_iff.mpr (λ x _, rfl) +@[simp] +theorem map_const (l : list α) (b : β) : map (function.const α b) l = replicate l.length b := +map_eq_replicate_iff.mpr (λ x _, rfl) theorem eq_of_mem_map_const {b₁ b₂ : β} {l : list α} (h : b₁ ∈ map (function.const α b₂) l) : b₁ = b₂ := -by rw map_const at h; exact eq_of_mem_repeat h +by rw map_const at h; exact eq_of_mem_replicate h /-! ### map₂ -/ @@ -1865,10 +1856,10 @@ theorem take_take : ∀ (n m) (l : list α), take n (take m l) = take (min n m) | (succ n) (succ m) nil := by simp only [take_nil] | (succ n) (succ m) (a::l) := by simp only [take, min_succ_succ, take_take n m l]; split; refl -theorem take_repeat (a : α) : ∀ (n m : ℕ), take n (repeat a m) = repeat a (min n m) +theorem take_replicate (a : α) : ∀ (n m : ℕ), take n (replicate m a) = replicate (min n m) a | n 0 := by simp | 0 m := by simp -| (succ n) (succ m) := by simp [min_succ_succ, take_repeat] +| (succ n) (succ m) := by simp [min_succ_succ, take_replicate] lemma map_take {α β : Type*} (f : α → β) : ∀ (L : list α) (i : ℕ), (L.take i).map f = (L.map f).take i @@ -2188,7 +2179,7 @@ variable [inhabited α] | 0 l := rfl | (n+1) l := congr_arg succ (take'_length _ _) -@[simp] theorem take'_nil : ∀ n, take' n (@nil α) = repeat default n +@[simp] theorem take'_nil : ∀ n, take' n (@nil α) = replicate n default | 0 := rfl | (n+1) := congr_arg (cons _) (take'_nil _) @@ -4140,7 +4131,7 @@ def decidable_nthd_nil_ne {α} (a : α) : decidable_pred @[simp] lemma nthd_singleton_default_eq (n : ℕ) : [d].nthd d n = d := by { cases n; simp } -@[simp] lemma nthd_repeat_default_eq (r n : ℕ) : (repeat d r).nthd d n = d := +@[simp] lemma nthd_replicate_default_eq (r n : ℕ) : (replicate r d).nthd d n = d := begin induction r with r IH generalizing n, { simp }, diff --git a/src/data/list/big_operators/basic.lean b/src/data/list/big_operators/basic.lean index 738d87ea3f9e6..b17f6051f0435 100644 --- a/src/data/list/big_operators/basic.lean +++ b/src/data/list/big_operators/basic.lean @@ -51,17 +51,17 @@ lemma prod_eq_foldr : l.prod = foldr (*) 1 l := list.rec_on l rfl $ λ a l ihl, by rw [prod_cons, foldr_cons, ihl] @[simp, priority 500, to_additive] -theorem prod_repeat (a : M) (n : ℕ) : (repeat a n).prod = a ^ n := +theorem prod_replicate (a : M) (n : ℕ) : (replicate n a).prod = a ^ n := begin induction n with n ih, { rw pow_zero, refl }, - { rw [list.repeat_succ, list.prod_cons, ih, pow_succ] } + { rw [list.replicate_succ, list.prod_cons, ih, pow_succ] } end @[to_additive sum_eq_card_nsmul] lemma prod_eq_pow_card (l : list M) (m : M) (h : ∀ (x ∈ l), x = m) : l.prod = m ^ l.length := -by rw [← prod_repeat, ← list.eq_repeat.mpr ⟨rfl, h⟩] +by rw [← prod_replicate, ← eq_replicate_length.2 h] @[to_additive] lemma prod_hom_rel (l : list ι) {r : M → N → Prop} {f : ι → M} {g : ι → N} (h₁ : r 1 1) @@ -97,8 +97,8 @@ lemma prod_map_neg {α} [comm_monoid α] [has_distrib_neg α] (l : list α) : begin convert @prod_map_mul α α _ l (λ _, -1) id, { ext, rw neg_one_mul, refl }, - { convert (prod_repeat _ _).symm, rw eq_repeat, - use l.length_map _, intro, rw mem_map, rintro ⟨_, _, rfl⟩, refl }, + { rw [← prod_replicate, map_eq_replicate_iff.2], + exact λ _ _, rfl }, { rw l.map_id }, end @@ -268,7 +268,7 @@ lemma prod_le_pow_card [preorder M] [covariant_class M M (function.swap (*)) (≤)] [covariant_class M M (*) (≤)] (l : list M) (n : M) (h : ∀ (x ∈ l), x ≤ n) : l.prod ≤ n ^ l.length := -by simpa only [map_id'', map_const, prod_repeat] using prod_le_prod' h +by simpa only [map_id'', map_const, prod_replicate] using prod_le_prod' 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 ι} @@ -478,8 +478,8 @@ lemma prod_map_erase [decidable_eq ι] [comm_monoid M] (f : ι → M) {a} : mul_left_comm (f a) (f b)], } end -@[simp] lemma sum_const_nat (m n : ℕ) : sum (list.repeat m n) = m * n := -by induction n; [refl, simp only [*, repeat_succ, sum_cons, nat.mul_succ, add_comm]] +@[simp] lemma sum_const_nat (m n : ℕ) : sum (replicate m n) = m * n := +by rw [sum_replicate, smul_eq_mul] /-- The product of a list of positive natural numbers is positive, and likewise for any nontrivial ordered semiring. -/ diff --git a/src/data/list/big_operators/lemmas.lean b/src/data/list/big_operators/lemmas.lean index 520d2cbba6d99..5cb2b62d1593f 100644 --- a/src/data/list/big_operators/lemmas.lean +++ b/src/data/list/big_operators/lemmas.lean @@ -55,7 +55,7 @@ lemma pow_card_le_prod [monoid M] [preorder M] @[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 _), - λ h, by rw [eq_repeat.2 ⟨rfl, h⟩, prod_repeat, one_pow]⟩ + λ h, by rw [eq_replicate.2 ⟨rfl, h⟩, prod_replicate, one_pow]⟩ /-- If a product of integers is `-1`, then at least one factor must be `-1`. -/ lemma neg_one_mem_of_prod_eq_neg_one {l : list ℤ} (h : l.prod = -1) : (-1 : ℤ) ∈ l := diff --git a/src/data/list/count.lean b/src/data/list/count.lean index a26f76e73b4c8..2f87dbeb2ae34 100644 --- a/src/data/list/count.lean +++ b/src/data/list/count.lean @@ -164,27 +164,27 @@ lemma not_mem_of_count_eq_zero {a : α} {l : list α} (h : count a l = 0) : a @[simp] lemma count_eq_length {a : α} {l} : count a l = l.length ↔ ∀ b ∈ l, a = b := countp_eq_length _ -@[simp] lemma count_repeat_self (a : α) (n : ℕ) : count a (repeat a n) = n := -by rw [count, countp_eq_length_filter, filter_eq_self.2, length_repeat]; - exact λ b m, (eq_of_mem_repeat m).symm +@[simp] lemma count_replicate_self (a : α) (n : ℕ) : count a (replicate n a) = n := +by rw [count, countp_eq_length_filter, filter_eq_self.2, length_replicate]; + exact λ b m, (eq_of_mem_replicate m).symm -lemma count_repeat (a b : α) (n : ℕ) : count a (repeat b n) = if a = b then n else 0 := +lemma count_replicate (a b : α) (n : ℕ) : count a (replicate n b) = if a = b then n else 0 := begin split_ifs with h, - exacts [h ▸ count_repeat_self _ _, count_eq_zero_of_not_mem (mt eq_of_mem_repeat h)] + exacts [h ▸ count_replicate_self _ _, count_eq_zero_of_not_mem $ mt eq_of_mem_replicate h] end -lemma le_count_iff_repeat_sublist {a : α} {l : list α} {n : ℕ} : - n ≤ count a l ↔ repeat a n <+ l := -⟨λ h, ((repeat_sublist_repeat a).2 h).trans $ - have filter (eq a) l = repeat a (count a l), from eq_repeat.2 +lemma le_count_iff_replicate_sublist {a : α} {l : list α} {n : ℕ} : + n ≤ count a l ↔ replicate n a <+ l := +⟨λ h, ((replicate_sublist_replicate a).2 h).trans $ + have filter (eq a) l = replicate (count a l) a, from eq_replicate.2 ⟨by simp only [count, countp_eq_length_filter], λ b m, (of_mem_filter m).symm⟩, by rw ← this; apply filter_sublist, - λ h, by simpa only [count_repeat_self] using h.count_le a⟩ + λ h, by simpa only [count_replicate_self] using h.count_le a⟩ -lemma repeat_count_eq_of_count_eq_length {a : α} {l : list α} (h : count a l = length l) : - repeat a (count a l) = l := -(le_count_iff_repeat_sublist.mp le_rfl).eq_of_length $ (length_repeat a (count a l)).trans h +lemma replicate_count_eq_of_count_eq_length {a : α} {l : list α} (h : count a l = length l) : + replicate (count a l) a = l := +(le_count_iff_replicate_sublist.mp le_rfl).eq_of_length $ (length_replicate (count a l) a).trans h @[simp] lemma count_filter {p} [decidable_pred p] {a} {l : list α} (h : p a) : count a (filter p l) = count a l := diff --git a/src/data/list/dedup.lean b/src/data/list/dedup.lean index b65f1122ff147..2bfdd74a3ad16 100644 --- a/src/data/list/dedup.lean +++ b/src/data/list/dedup.lean @@ -74,11 +74,11 @@ begin rw [dedup_cons_of_not_mem' h, insert_of_not_mem h]] end -lemma repeat_dedup {x : α} : ∀ {k}, k ≠ 0 → (repeat x k).dedup = [x] +lemma replicate_dedup {x : α} : ∀ {k}, k ≠ 0 → (replicate k x).dedup = [x] | 0 h := (h rfl).elim | 1 _ := rfl -| (n+2) _ := by rw [repeat_succ, dedup_cons_of_mem (mem_repeat.2 ⟨n.succ_ne_zero, rfl⟩), - repeat_dedup n.succ_ne_zero] +| (n+2) _ := by rw [replicate_succ, dedup_cons_of_mem (mem_replicate.2 ⟨n.succ_ne_zero, rfl⟩), + replicate_dedup n.succ_ne_zero] lemma count_dedup (l : list α) (a : α) : l.dedup.count a = if a ∈ l then 1 else 0 := diff --git a/src/data/list/defs.lean b/src/data/list/defs.lean index b20ef680b6575..d548a377c57b7 100644 --- a/src/data/list/defs.lean +++ b/src/data/list/defs.lean @@ -25,6 +25,11 @@ variables {α β γ δ ε ζ : Type*} instance [decidable_eq α] : has_sdiff (list α) := ⟨ list.diff ⟩ +/-- Create a list of `n` copies of `a`. Same as `function.swap list.repeat`. -/ +@[simp] def replicate : ℕ → α → list α +| 0 _ := [] +| (succ n) a := a :: replicate n a + /-- Split a list at an index. split_at 2 [a, b, c] = ([a, b], [c]) -/ diff --git a/src/data/list/duplicate.lean b/src/data/list/duplicate.lean index efa5f35a83d93..e142708891e1e 100644 --- a/src/data/list/duplicate.lean +++ b/src/data/list/duplicate.lean @@ -129,7 +129,7 @@ lemma duplicate.not_nodup (h : x ∈+ l) : ¬ nodup l := λ H, nodup_iff_forall_not_duplicate.mp H _ h lemma duplicate_iff_two_le_count [decidable_eq α] : (x ∈+ l) ↔ 2 ≤ count x l := -by simp [duplicate_iff_sublist, le_count_iff_repeat_sublist] +by simp [duplicate_iff_sublist, le_count_iff_replicate_sublist] instance decidable_duplicate [decidable_eq α] (x : α) : ∀ (l : list α), decidable (x ∈+ l) | [] := is_false (not_duplicate_nil x) diff --git a/src/data/list/nodup.lean b/src/data/list/nodup.lean index fc3b413305d01..1b0a1b4365216 100644 --- a/src/data/list/nodup.lean +++ b/src/data/list/nodup.lean @@ -125,14 +125,14 @@ index_of_nth_le $ index_of_lt_length.2 $ nth_le_mem _ _ _ theorem nodup_iff_count_le_one [decidable_eq α] {l : list α} : nodup l ↔ ∀ a, count a l ≤ 1 := nodup_iff_sublist.trans $ forall_congr $ λ a, -have [a, a] <+ l ↔ 1 < count a l, from (@le_count_iff_repeat_sublist _ _ a l 2).symm, +have [a, a] <+ l ↔ 1 < count a l, from (@le_count_iff_replicate_sublist _ _ a l 2).symm, (not_congr this).trans not_lt -theorem nodup_repeat (a : α) : ∀ {n : ℕ}, nodup (repeat a n) ↔ n ≤ 1 +theorem nodup_replicate (a : α) : ∀ {n : ℕ}, nodup (replicate n a) ↔ n ≤ 1 | 0 := by simp [nat.zero_le] | 1 := by simp | (n+2) := iff_of_false - (λ H, nodup_iff_sublist.1 H a ((repeat_sublist_repeat _).2 (nat.le_add_left 2 n))) + (λ H, nodup_iff_sublist.1 H a ((replicate_sublist_replicate _).2 (nat.le_add_left 2 n))) (not_le_of_lt $ nat.le_add_left 2 n) @[simp] theorem count_eq_one_of_mem [decidable_eq α] {a : α} {l : list α} diff --git a/src/data/list/nodup_equiv_fin.lean b/src/data/list/nodup_equiv_fin.lean index f423ffcf37b06..052911bb4085a 100644 --- a/src/data/list/nodup_equiv_fin.lean +++ b/src/data/list/nodup_equiv_fin.lean @@ -201,7 +201,7 @@ lemma duplicate_iff_exists_distinct_nth_le {l : list α} {x : α} : x = l.nth_le n hn ∧ x = l.nth_le m hm := begin classical, - rw [duplicate_iff_two_le_count, le_count_iff_repeat_sublist, + rw [duplicate_iff_two_le_count, le_count_iff_replicate_sublist, sublist_iff_exists_fin_order_embedding_nth_le_eq], split, { rintro ⟨f, hf⟩, @@ -214,7 +214,7 @@ begin { simp }, { simp [hnm] }, { simp }, - { simp only [nat.lt_succ_iff, nat.succ_le_succ_iff, repeat, length, nonpos_iff_eq_zero] + { simp only [nat.lt_succ_iff, nat.succ_le_succ_iff, replicate, length, nonpos_iff_eq_zero] at hi hj, simp [hi, hj] } }, { rintros ⟨⟨_|i⟩, hi⟩, diff --git a/src/data/list/of_fn.lean b/src/data/list/of_fn.lean index f131ed93579ed..18f556758be76 100644 --- a/src/data/list/of_fn.lean +++ b/src/data/list/of_fn.lean @@ -170,7 +170,7 @@ end by simp only [mem_of_fn, set.forall_range_iff] @[simp] lemma of_fn_const (n : ℕ) (c : α) : - of_fn (λ i : fin n, c) = repeat c n := + of_fn (λ i : fin n, c) = replicate n c := nat.rec_on n (by simp) $ λ n ihn, by simp [ihn] /-- Lists are equivalent to the sigma type of tuples of a given length. -/ diff --git a/src/data/list/pairwise.lean b/src/data/list/pairwise.lean index e702acbb2c163..b186482bb8485 100644 --- a/src/data/list/pairwise.lean +++ b/src/data/list/pairwise.lean @@ -313,10 +313,10 @@ theorem pairwise_iff_nth_le {R} : ∀ {l : list α}, exact H _ _ (succ_lt_succ h) (succ_pos _) } end -lemma pairwise_repeat {α : Type*} {r : α → α → Prop} {x : α} (hx : r x x) : - ∀ (n : ℕ), pairwise r (repeat x n) +lemma pairwise_replicate {α : Type*} {r : α → α → Prop} {x : α} (hx : r x x) : + ∀ (n : ℕ), pairwise r (replicate n x) | 0 := by simp -| (n+1) := by simp [hx, mem_repeat, pairwise_repeat n] +| (n+1) := by simp [hx, mem_replicate, pairwise_replicate n] /-! ### Pairwise filtering -/ diff --git a/src/data/list/perm.lean b/src/data/list/perm.lean index 5064dbc38934e..557cce5f96c3a 100644 --- a/src/data/list/perm.lean +++ b/src/data/list/perm.lean @@ -150,26 +150,24 @@ theorem perm_cons_append_cons {l l₁ l₂ : list α} (a : α) (p : l ~ l₁++l a::l ~ l₁++(a::l₂) := (p.cons a).trans perm_middle.symm -@[simp] theorem perm_repeat {a : α} {n : ℕ} {l : list α} : l ~ repeat a n ↔ l = repeat a n := -⟨λ p, (eq_repeat.2 - ⟨p.length_eq.trans $ length_repeat _ _, - λ b m, eq_of_mem_repeat $ p.subset m⟩), - λ h, h ▸ perm.refl _⟩ +@[simp] theorem perm_replicate {a : α} {n : ℕ} {l : list α} : + l ~ replicate n a ↔ l = replicate n a := +⟨λ p, eq_replicate.2 + ⟨p.length_eq.trans $ length_replicate _ _, λ b m, eq_of_mem_replicate $ p.subset m⟩, + λ h, h ▸ perm.refl _⟩ -@[simp] theorem repeat_perm {a : α} {n : ℕ} {l : list α} : repeat a n ~ l ↔ repeat a n = l := -(perm_comm.trans perm_repeat).trans eq_comm +@[simp] theorem replicate_perm {a : α} {n : ℕ} {l : list α} : + replicate n a ~ l ↔ replicate n a = l := +(perm_comm.trans perm_replicate).trans eq_comm @[simp] theorem perm_singleton {a : α} {l : list α} : l ~ [a] ↔ l = [a] := -@perm_repeat α a 1 l +@perm_replicate α a 1 l @[simp] theorem singleton_perm {a : α} {l : list α} : [a] ~ l ↔ [a] = l := -@repeat_perm α a 1 l +@replicate_perm α a 1 l -theorem perm.eq_singleton {a : α} {l : list α} (p : l ~ [a]) : l = [a] := -perm_singleton.1 p - -theorem perm.singleton_eq {a : α} {l : list α} (p : [a] ~ l) : [a] = l := -p.symm.eq_singleton.symm +alias perm_singleton ↔ perm.eq_singleton _ +alias singleton_perm ↔ perm.singleton_eq _ theorem singleton_perm_singleton {a b : α} : [a] ~ [b] ↔ a = b := by simp @@ -736,12 +734,12 @@ theorem perm_iff_count {l₁ l₂ : list α} : l₁ ~ l₂ ↔ ∀ a, count a l by_cases b = a; simp [h] at H ⊢; assumption } end⟩ -theorem perm_repeat_append_repeat {l : list α} {a b : α} {m n : ℕ} (h : a ≠ b) : - l ~ repeat a m ++ repeat b n ↔ count a l = m ∧ count b l = n ∧ l ⊆ [a, b] := +theorem perm_replicate_append_replicate {l : list α} {a b : α} {m n : ℕ} (h : a ≠ b) : + l ~ replicate m a ++ replicate n b ↔ count a l = m ∧ count b l = n ∧ l ⊆ [a, b] := begin rw [perm_iff_count, ← decidable.and_forall_ne a, ← decidable.and_forall_ne b], suffices : l ⊆ [a, b] ↔ ∀ c, c ≠ b → c ≠ a → c ∉ l, - { simp [count_repeat, h, h.symm, this] { contextual := tt } }, + { simp [count_replicate, h, h.symm, this] { contextual := tt } }, simp_rw [ne.def, ← and_imp, ← not_or_distrib, decidable.not_imp_not, subset_def, mem_cons_iff, not_mem_nil, or_false, or_comm], end diff --git a/src/data/list/rotate.lean b/src/data/list/rotate.lean index 17fbef984564c..9a8d9f7e1f798 100644 --- a/src/data/list/rotate.lean +++ b/src/data/list/rotate.lean @@ -99,10 +99,9 @@ by rw [rotate_eq_rotate', rotate_eq_rotate', rotate'_cons_succ] @[simp] lemma length_rotate (l : list α) (n : ℕ) : (l.rotate n).length = l.length := by rw [rotate_eq_rotate', length_rotate'] -lemma rotate_repeat (a : α) (n : ℕ) (k : ℕ) : - (repeat a n).rotate k = repeat a n := -eq_repeat.2 ⟨by rw [length_rotate, length_repeat], - λ b hb, eq_of_mem_repeat $ mem_rotate.1 hb⟩ +lemma rotate_replicate (a : α) (n : ℕ) (k : ℕ) : (replicate n a).rotate k = replicate n a := +eq_replicate.2 ⟨by rw [length_rotate, length_replicate], + λ b hb, eq_of_mem_replicate $ mem_rotate.1 hb⟩ lemma rotate_eq_drop_append_take {l : list α} {n : ℕ} : n ≤ l.length → l.rotate n = l.drop n ++ l.take n := @@ -170,7 +169,7 @@ by rw [eq_comm, rotate_eq_nil_iff, eq_comm] @[simp] lemma rotate_singleton (x : α) (n : ℕ) : [x].rotate n = [x] := -rotate_repeat x 1 n +rotate_replicate x 1 n lemma zip_with_rotate_distrib {α β γ : Type*} (f : α → β → γ) (l : list α) (l' : list β) (n : ℕ) (h : l.length = l'.length) : @@ -241,22 +240,22 @@ lemma head'_rotate {l : list α} {n : ℕ} (h : n < l.length) : head' (l.rotate n) = l.nth n := by rw [← nth_zero, nth_rotate (n.zero_le.trans_lt h), zero_add, nat.mod_eq_of_lt h] -lemma rotate_eq_self_iff_eq_repeat [hα : nonempty α] : - ∀ {l : list α}, (∀ n, l.rotate n = l) ↔ ∃ a, l = repeat a l.length +lemma rotate_eq_self_iff_eq_replicate [hα : nonempty α] : + ∀ {l : list α}, (∀ n, l.rotate n = l) ↔ ∃ a, l = replicate l.length a | [] := by simp -| (a :: l) := ⟨λ h, ⟨a, ext_le (length_repeat _ _).symm $ λ n h₁ h₂, +| (a :: l) := ⟨λ h, ⟨a, ext_le (length_replicate _ _).symm $ λ n h₁ h₂, begin inhabit α, - rw [nth_le_repeat, ← option.some_inj, ← nth_le_nth, ← head'_rotate h₁, h, head'] - end⟩, λ ⟨b, hb⟩ n, by rw [hb, rotate_repeat]⟩ + rw [nth_le_replicate, ← option.some_inj, ← nth_le_nth, ← head'_rotate h₁, h, head'] + end⟩, λ ⟨b, hb⟩ n, by rw [hb, rotate_replicate]⟩ -lemma rotate_one_eq_self_iff_eq_repeat [nonempty α] {l : list α} : - l.rotate 1 = l ↔ ∃ a : α, l = repeat a l.length := -⟨λ h, rotate_eq_self_iff_eq_repeat.mp (λ n, nat.rec l.rotate_zero +lemma rotate_one_eq_self_iff_eq_replicate [nonempty α] {l : list α} : + l.rotate 1 = l ↔ ∃ a : α, l = list.replicate l.length a := +⟨λ h, rotate_eq_self_iff_eq_replicate.mp (λ n, nat.rec l.rotate_zero (λ n hn, by rwa [nat.succ_eq_add_one, ←l.rotate_rotate, hn]) n), - λ h, rotate_eq_self_iff_eq_repeat.mpr h 1⟩ + λ h, rotate_eq_self_iff_eq_replicate.mpr h 1⟩ -lemma rotate_injective (n : ℕ) : injective (λ l : list α, l.rotate n) := +lemma rotate_injective (n : ℕ) : function.injective (λ l : list α, l.rotate n) := begin rintro l₁ l₂ (h : l₁.rotate n = l₂.rotate n), have : length l₁ = length l₂, by simpa only [length_rotate] using congr_arg length h, @@ -282,7 +281,7 @@ begin end @[simp] lemma rotate_eq_singleton_iff {l : list α} {n : ℕ} {x : α} : l.rotate n = [x] ↔ l = [x] := -⟨λ h, by rw [rotate_eq_iff.1 h, rotate_singleton], λ h, h.symm ▸ rotate_singleton _ _⟩ +by rw [rotate_eq_iff, rotate_singleton] @[simp] lemma singleton_eq_rotate_iff {l : list α} {n : ℕ} {x : α} : [x] = l.rotate n ↔ [x] = l := by rw [eq_comm, rotate_eq_singleton_iff, eq_comm] diff --git a/src/data/list/sigma.lean b/src/data/list/sigma.lean index 46a0917c60e95..600bbc18f7853 100644 --- a/src/data/list/sigma.lean +++ b/src/data/list/sigma.lean @@ -262,7 +262,7 @@ theorem lookup_all_sublist (a : α) : theorem lookup_all_length_le_one (a : α) {l : list (sigma β)} (h : l.nodupkeys) : length (lookup_all a l) ≤ 1 := by have := nodup.sublist ((lookup_all_sublist a l).map _) h; - rw map_map at this; rwa [← nodup_repeat, ← map_const _ a] + rw map_map at this; rwa [← nodup_replicate, ← map_const _ a] theorem lookup_all_eq_lookup (a : α) {l : list (sigma β)} (h : l.nodupkeys) : lookup_all a l = (lookup a l).to_list := diff --git a/src/data/list/sort.lean b/src/data/list/sort.lean index 86373474b8734..68fecf73bf0e4 100644 --- a/src/data/list/sort.lean +++ b/src/data/list/sort.lean @@ -67,8 +67,8 @@ begin have : ∀ (x : α) (h : x ∈ u₂), x = a := λ x m, antisymm ((pairwise_append.1 s₂).2.2 _ m a (mem_cons_self _ _)) (h₁ _ (by simp [m])), - rw [(@eq_repeat _ a (length u₂ + 1) (a::u₂)).2, - (@eq_repeat _ a (length u₂ + 1) (u₂++[a])).2]; + rw [(@eq_replicate _ a (length u₂ + 1) (a::u₂)).2, + (@eq_replicate _ a (length u₂ + 1) (u₂++[a])).2]; split; simp [iff_true_intro this, or_comm] } end diff --git a/src/data/multiset/basic.lean b/src/data/multiset/basic.lean index d6f6ab88f8906..a29f9e74da07e 100644 --- a/src/data/multiset/basic.lean +++ b/src/data/multiset/basic.lean @@ -618,94 +618,88 @@ subrelation.wf (λ _ _, multiset.card_lt_of_lt) (measure_wf multiset.card) instance is_well_founded_lt : _root_.well_founded_lt (multiset α) := ⟨well_founded_lt⟩ -/-! ### `multiset.repeat` -/ +/-! ### `multiset.replicate` -/ -/-- `repeat a n` is the multiset containing only `a` with multiplicity `n`. -/ -def repeat (a : α) (n : ℕ) : multiset α := repeat a n +/-- `replicate n a` is the multiset containing only `a` with multiplicity `n`. -/ +def replicate (n : ℕ) (a : α) : multiset α := replicate n a -lemma coe_repeat (a : α) (n : ℕ) : (list.repeat a n : multiset α) = repeat a n := rfl +lemma coe_replicate (a : α) (n : ℕ) : (list.replicate n a : multiset α) = replicate n a := rfl -@[simp] lemma repeat_zero (a : α) : repeat a 0 = 0 := rfl +@[simp] lemma replicate_zero (a : α) : replicate 0 a = 0 := rfl -@[simp] lemma repeat_succ (a : α) (n) : repeat a (n+1) = a ::ₘ repeat a n := by simp [repeat] +@[simp] lemma replicate_succ (a : α) (n) : replicate (n+1) a = a ::ₘ replicate n a := rfl -lemma repeat_add (a : α) (m n : ℕ) : repeat a (m + n) = repeat a m + repeat a n := -congr_arg _ $ list.repeat_add _ _ _ +lemma replicate_add (a : α) (m n : ℕ) : replicate (m + n) a = replicate m a + replicate n a := +congr_arg _ $ list.replicate_add _ _ _ -/-- `multiset.repeat` as an `add_monoid_hom`. -/ -@[simps] def repeat_add_monoid_hom (a : α) : ℕ →+ multiset α := -{ to_fun := repeat a, map_zero' := repeat_zero a, map_add' := repeat_add a } +/-- `multiset.replicate` as an `add_monoid_hom`. -/ +@[simps] def replicate_add_monoid_hom (a : α) : ℕ →+ multiset α := +{ to_fun := λ n, replicate n a, map_zero' := replicate_zero a, map_add' := replicate_add a } -@[simp] lemma repeat_one (a : α) : repeat a 1 = {a} := -by simp only [repeat_succ, ←cons_zero, eq_self_iff_true, repeat_zero, cons_inj_right] +@[simp] lemma replicate_one (a : α) : replicate 1 a = {a} := rfl -@[simp] lemma card_repeat : ∀ (a : α) n, card (repeat a n) = n := length_repeat +@[simp] lemma card_replicate : ∀ n (a : α), card (replicate n a) = n := length_replicate -lemma mem_repeat {a b : α} {n : ℕ} : b ∈ repeat a n ↔ n ≠ 0 ∧ b = a := mem_repeat +lemma mem_replicate {a b : α} {n : ℕ} : b ∈ replicate n a ↔ n ≠ 0 ∧ b = a := mem_replicate -theorem eq_of_mem_repeat {a b : α} {n} : b ∈ repeat a n → b = a := eq_of_mem_repeat +theorem eq_of_mem_replicate {a b : α} {n} : b ∈ replicate n a → b = a := eq_of_mem_replicate -theorem eq_repeat' {a : α} {s : multiset α} : s = repeat a s.card ↔ ∀ b ∈ s, b = a := -quot.induction_on s $ λ l, iff.trans ⟨λ h, - (perm_repeat.1 $ (quotient.exact h)), congr_arg coe⟩ eq_repeat' +theorem eq_replicate_card {a : α} {s : multiset α} : s = replicate s.card a ↔ ∀ b ∈ s, b = a := +quot.induction_on s $ λ l, coe_eq_coe.trans $ perm_replicate.trans $ eq_replicate_length -theorem eq_repeat_of_mem {a : α} {s : multiset α} : (∀ b ∈ s, b = a) → s = repeat a s.card := -eq_repeat'.2 +alias eq_replicate_card ↔ _ eq_replicate_of_mem -theorem eq_repeat {a : α} {n} {s : multiset α} : s = repeat a n ↔ card s = n ∧ ∀ b ∈ s, b = a := -⟨λ h, h.symm ▸ ⟨card_repeat _ _, λ b, eq_of_mem_repeat⟩, - λ ⟨e, al⟩, e ▸ eq_repeat_of_mem al⟩ +theorem eq_replicate {a : α} {n} {s : multiset α} : + s = replicate n a ↔ card s = n ∧ ∀ b ∈ s, b = a := +⟨λ h, h.symm ▸ ⟨card_replicate _ _, λ b, eq_of_mem_replicate⟩, + λ ⟨e, al⟩, e ▸ eq_replicate_of_mem al⟩ -lemma repeat_left_injective {n : ℕ} (hn : n ≠ 0) : function.injective (λ a : α, repeat a n) := -λ a b h, (eq_repeat.1 h).2 _ $ mem_repeat.2 ⟨hn, rfl⟩ +lemma replicate_right_injective {n : ℕ} (hn : n ≠ 0) : + function.injective (replicate n : α → multiset α) := +λ a b h, (eq_replicate.1 h).2 _ $ mem_replicate.2 ⟨hn, rfl⟩ -@[simp] lemma repeat_left_inj {a b : α} {n : ℕ} (h : n ≠ 0) : repeat a n = repeat b n ↔ a = b := -(repeat_left_injective h).eq_iff +@[simp] lemma replicate_right_inj {a b : α} {n : ℕ} (h : n ≠ 0) : + replicate n a = replicate n b ↔ a = b := +(replicate_right_injective h).eq_iff -theorem repeat_injective (a : α) : function.injective (repeat a) := -λ m n h, by rw [← (eq_repeat.1 h).1, card_repeat] +theorem replicate_left_injective (a : α) : function.injective (λ n, replicate n a) := +λ m n h, by rw [← (eq_replicate.1 h).1, card_replicate] -theorem repeat_subset_singleton : ∀ (a : α) n, repeat a n ⊆ {a} := repeat_subset_singleton +theorem replicate_subset_singleton : ∀ (a : α) n, replicate n a ⊆ {a} := replicate_subset_singleton -theorem repeat_le_coe {a : α} {n} {l : list α} : repeat a n ≤ l ↔ list.repeat a n <+ l := -⟨λ ⟨l', p, s⟩, (perm_repeat.1 p) ▸ s, sublist.subperm⟩ +theorem replicate_le_coe {a : α} {n} {l : list α} : + replicate n a ≤ l ↔ list.replicate n a <+ l := +⟨λ ⟨l', p, s⟩, (perm_replicate.1 p) ▸ s, sublist.subperm⟩ -theorem nsmul_singleton (a : α) (n) : n • ({a} : multiset α) = repeat a n := +theorem nsmul_singleton (a : α) (n) : n • ({a} : multiset α) = replicate n a := begin - refine eq_repeat.mpr ⟨_, λ b hb, mem_singleton.mp (mem_of_mem_nsmul hb)⟩, + refine eq_replicate.mpr ⟨_, λ b hb, mem_singleton.mp (mem_of_mem_nsmul hb)⟩, rw [card_nsmul, card_singleton, mul_one] end -lemma nsmul_repeat {a : α} (n m : ℕ) : n • (repeat a m) = repeat a (n * m) := -((repeat_add_monoid_hom a).map_nsmul _ _).symm +lemma nsmul_replicate {a : α} (n m : ℕ) : n • replicate m a = replicate (n * m) a := +((replicate_add_monoid_hom a).map_nsmul _ _).symm -lemma repeat_le_repeat (a : α) {k n : ℕ} : - repeat a k ≤ repeat a n ↔ k ≤ n := -trans (by rw [← repeat_le_coe, coe_repeat]) (list.repeat_sublist_repeat a) +lemma replicate_le_replicate (a : α) {k n : ℕ} : + replicate k a ≤ replicate n a ↔ k ≤ n := +replicate_le_coe.trans $ list.replicate_sublist_replicate _ -lemma le_repeat_iff {m : multiset α} {a : α} {n : ℕ} : - m ≤ repeat a n ↔ ∃ (k ≤ n), m = repeat a k := -quot.induction_on m (λ l, show (l : multiset α) ≤ repeat a n ↔ ∃ (k ≤ n), ↑l = repeat a k, -begin - simp only [← coe_repeat, coe_le, subperm, sublist_repeat_iff, coe_eq_coe, perm_repeat], - split, - { rintros ⟨l, hl, k, h, rfl⟩, - rw [perm_comm, perm_repeat] at hl, - exact ⟨k, h, hl⟩ }, - { rintros ⟨k, h, hl⟩, - exact ⟨l, refl _, k, h, hl⟩ } -end) - -lemma lt_repeat_succ {m : multiset α} {x : α} {n : ℕ} : - m < repeat x (n + 1) ↔ m ≤ repeat x n := +lemma le_replicate_iff {m : multiset α} {a : α} {n : ℕ} : + m ≤ replicate n a ↔ ∃ (k ≤ n), m = replicate k a := +⟨λ h, ⟨m.card, (card_mono h).trans_eq (card_replicate _ _), eq_replicate_card.2 $ + λ b hb, eq_of_mem_replicate $ subset_of_le h hb⟩, + λ ⟨k, hkn, hm⟩, hm.symm ▸ (replicate_le_replicate _).2 hkn⟩ + +lemma lt_replicate_succ {m : multiset α} {x : α} {n : ℕ} : + m < replicate (n + 1) x ↔ m ≤ replicate n x := begin rw lt_iff_cons_le, split, { rintros ⟨x', hx'⟩, - have := eq_of_mem_repeat (mem_of_le hx' (mem_cons_self _ _)), - rwa [this, repeat_succ, cons_le_cons_iff] at hx' }, + have := eq_of_mem_replicate (mem_of_le hx' (mem_cons_self _ _)), + rwa [this, replicate_succ, cons_le_cons_iff] at hx' }, { intro h, - rw repeat_succ, + rw replicate_succ, exact ⟨x, cons_le_cons _ h⟩ } end @@ -862,8 +856,9 @@ by { ext, simp } @[simp] theorem map_singleton (f : α → β) (a : α) : ({a} : multiset α).map f = {f a} := rfl -theorem map_repeat (f : α → β) (a : α) (k : ℕ) : (repeat a k).map f = repeat (f a) k := by -{ induction k, simp, simpa } +@[simp] theorem map_replicate (f : α → β) (a : α) (k : ℕ) : + (replicate k a).map f = replicate k (f a) := +by simp only [← coe_replicate, coe_map, map_replicate] @[simp] theorem map_add (f : α → β) (s t) : map f (s + t) = map f s + map f t := quotient.induction_on₂ s t $ λ l₁ l₂, congr_arg coe $ map_append _ _ _ @@ -941,12 +936,13 @@ quot.induction_on s $ λ l, congr_arg coe $ map_id _ @[simp] lemma map_id' (s : multiset α) : map (λx, x) s = s := map_id s -@[simp] theorem map_const (s : multiset α) (b : β) : map (function.const α b) s = repeat b s.card := +@[simp] theorem map_const (s : multiset α) (b : β) : + map (function.const α b) s = replicate s.card b := quot.induction_on s $ λ l, congr_arg coe $ map_const _ _ theorem eq_of_mem_map_const {b₁ b₂ : β} {l : list α} (h : b₁ ∈ map (function.const α b₂) l) : b₁ = b₂ := -eq_of_mem_repeat $ by rwa map_const at h +eq_of_mem_replicate $ by rwa map_const at h @[simp] theorem map_le_map {f : α → β} {s t : multiset α} (h : s ≤ t) : map f s ≤ map f t := le_induction_on h $ λ l₁ l₂ h, (h.map f).subperm @@ -1794,29 +1790,20 @@ by simp [ne.def, count_eq_zero] theorem count_eq_card {a : α} {s} : count a s = card s ↔ ∀ (x ∈ s), a = x := countp_eq_card -@[simp] theorem count_repeat_self (a : α) (n : ℕ) : count a (repeat a n) = n := -by simp [repeat] +@[simp] theorem count_replicate_self (a : α) (n : ℕ) : count a (replicate n a) = n := +count_replicate_self _ _ -theorem count_repeat (a b : α) (n : ℕ) : - count a (repeat b n) = if (a = b) then n else 0 := -count_repeat a b n +theorem count_replicate (a b : α) (n : ℕ) : + count a (replicate n b) = if (a = b) then n else 0 := +count_replicate _ _ _ @[simp] theorem count_erase_self (a : α) (s : multiset α) : count a (erase s a) = pred (count a s) := -begin - by_cases a ∈ s, - { rw [(by rw cons_erase h : count a s = count a (a ::ₘ erase s a)), - count_cons_self]; refl }, - { rw [erase_of_not_mem h, count_eq_zero.2 h]; refl } -end +quotient.induction_on s $ count_erase_self a @[simp, priority 980] theorem count_erase_of_ne {a b : α} (ab : a ≠ b) (s : multiset α) : count a (erase s b) = count a s := -begin - by_cases b ∈ s, - { rw [← count_cons_of_ne ab, cons_erase h] }, - { rw [erase_of_not_mem h] } -end +quotient.induction_on s $ count_erase_of_ne ab @[simp] theorem count_sub (a : α) (s t : multiset α) : count a (s - t) = count a s - count a t := begin @@ -1838,8 +1825,9 @@ begin rw [← count_add, sub_add_inter, count_sub, tsub_add_min], end -theorem le_count_iff_repeat_le {a : α} {s : multiset α} {n : ℕ} : n ≤ count a s ↔ repeat a n ≤ s := -quot.induction_on s $ λ l, le_count_iff_repeat_sublist.trans repeat_le_coe.symm +theorem le_count_iff_replicate_le {a : α} {s : multiset α} {n : ℕ} : + n ≤ count a s ↔ replicate n a ≤ s := +quot.induction_on s $ λ l, le_count_iff_replicate_sublist.trans replicate_le_coe.symm @[simp] theorem count_filter_of_pos {p} [decidable_pred p] {a} {s : multiset α} (h : p a) : count a (filter p s) = count a s := @@ -1878,16 +1866,6 @@ instance : distrib_lattice (multiset α) := multiset.count_inter, multiset.sup_eq_union, multiset.count_union, multiset.inf_eq_inter], ..multiset.lattice } -theorem repeat_inf (s : multiset α) (a : α) (n : ℕ) : - (repeat a n) ⊓ s = repeat a (min (s.count a) n) := -begin - ext x, - rw [inf_eq_inter, count_inter, count_repeat, count_repeat], - by_cases x = a, - simp only [min_comm, h, if_true, eq_self_iff_true], - simp only [h, if_false, zero_min], -end - theorem count_map {α β : Type*} (f : α → β) (s : multiset α) [decidable_eq β] (b : β) : count b (map f s) = (s.filter (λ a, b = f a)).card := countp_map _ _ _ @@ -1900,8 +1878,8 @@ begin suffices : (filter (λ (a : α), f x = f a) s).count x = card (filter (λ (a : α), f x = f a) s), { rw [count, countp_map, ← this], exact count_filter_of_pos rfl }, - { rw eq_repeat.2 ⟨rfl, λ b hb, eq_comm.1 ((hf H (mem_filter.1 hb).left) (mem_filter.1 hb).right)⟩, - simp only [count_repeat, eq_self_iff_true, if_true, card_repeat]}, + { rw [eq_replicate_card.2 (λ b hb, ((hf H (mem_filter.1 hb).left) (mem_filter.1 hb).2).symm), + count_replicate_self, card_replicate] } end /-- `multiset.map f` preserves `count` if `f` is injective -/ @@ -1923,31 +1901,29 @@ calc m.attach.count a (multiset.count_map_eq_count' _ _ subtype.coe_injective _).symm ... = m.count (a : α) : congr_arg _ m.attach_map_coe -lemma filter_eq' (s : multiset α) (b : α) : s.filter (= b) = repeat b (count b s) := +lemma filter_eq' (s : multiset α) (b : α) : s.filter (= b) = replicate (count b s) b := begin ext a, - rw [count_repeat, count_filter], + rw [count_replicate, count_filter], exact if_ctx_congr iff.rfl (λ h, congr_arg _ h) (λ h, rfl), end -lemma filter_eq (s : multiset α) (b : α) : s.filter (eq b) = repeat b (count b s) := +lemma filter_eq (s : multiset α) (b : α) : s.filter (eq b) = replicate (count b s) b := by simp_rw [←filter_eq', eq_comm] -@[simp] lemma repeat_inter (x : α) (n : ℕ) (s : multiset α) : - repeat x n ∩ s = repeat x (min n (s.count x)) := +@[simp] lemma replicate_inter (x : α) (n : ℕ) (s : multiset α) : + replicate n x ∩ s = replicate (min n (s.count x)) x := begin - refine le_antisymm _ _, - { simp only [le_iff_count, count_inter, count_repeat], - intro a, - split_ifs with h, - { rw h }, - { rw [nat.zero_min] } }, - simp only [le_inter_iff, ← le_count_iff_repeat_le, count_inter, count_repeat_self], + ext y, + rw [count_inter, count_replicate, count_replicate], + by_cases y = x, + { simp only [h, if_pos rfl] }, + { simp only [h, if_false, zero_min] } end -@[simp] lemma inter_repeat (s : multiset α) (x : α) (n : ℕ) : - s ∩ repeat x n = repeat x (min (s.count x) n) := -by rw [inter_comm, repeat_inter, min_comm] +@[simp] lemma inter_replicate (s : multiset α) (x : α) (n : ℕ) : + s ∩ replicate n x = replicate (min (s.count x) n) x := +by rw [inter_comm, replicate_inter, min_comm] end @@ -2148,18 +2124,18 @@ begin { simpa using hc } } end -lemma rel_repeat_left {m : multiset α} {a : α} {r : α → α → Prop} {n : ℕ} : - (repeat a n).rel r m ↔ m.card = n ∧ ∀ x, x ∈ m → r a x := -⟨λ h, ⟨(card_eq_card_of_rel h).symm.trans (card_repeat _ _), λ x hx, begin +lemma rel_replicate_left {m : multiset α} {a : α} {r : α → α → Prop} {n : ℕ} : + (replicate n a).rel r m ↔ m.card = n ∧ ∀ x, x ∈ m → r a x := +⟨λ h, ⟨(card_eq_card_of_rel h).symm.trans (card_replicate _ _), λ x hx, begin obtain ⟨b, hb1, hb2⟩ := exists_mem_of_rel_of_mem (rel_flip.2 h) hx, - rwa eq_of_mem_repeat hb1 at hb2, + rwa eq_of_mem_replicate hb1 at hb2, end⟩, - λ h, rel_of_forall (λ x y hx hy, (eq_of_mem_repeat hx).symm ▸ (h.2 _ hy)) - (eq.trans (card_repeat _ _) h.1.symm)⟩ + λ h, rel_of_forall (λ x y hx hy, (eq_of_mem_replicate hx).symm ▸ (h.2 _ hy)) + (eq.trans (card_replicate _ _) h.1.symm)⟩ -lemma rel_repeat_right {m : multiset α} {a : α} {r : α → α → Prop} {n : ℕ} : - m.rel r (repeat a n) ↔ m.card = n ∧ ∀ x, x ∈ m → r x a := -by { rw [← rel_flip], exact rel_repeat_left } +lemma rel_replicate_right {m : multiset α} {a : α} {r : α → α → Prop} {n : ℕ} : + m.rel r (replicate n a) ↔ m.card = n ∧ ∀ x, x ∈ m → r x a := +rel_flip.trans rel_replicate_left lemma rel.trans (r : α → α → Prop) [is_trans α r] {s t u : multiset α} (r1 : rel r s t) (r2 : rel r t u) : diff --git a/src/data/multiset/bind.lean b/src/data/multiset/bind.lean index 031bc820e400e..42598b29f4b28 100644 --- a/src/data/multiset/bind.lean +++ b/src/data/multiset/bind.lean @@ -179,8 +179,7 @@ multiset.induction_on s (λ t u, rfl) $ λ a s IH t u, @[simp] lemma mem_product {s t} : ∀ {p : α × β}, p ∈ @product α β s t ↔ p.1 ∈ s ∧ p.2 ∈ t | (a, b) := by simp [product, and.left_comm] -@[simp] lemma card_product : (s ×ˢ t).card = s.card * t.card := -by simp [product, repeat, (∘), mul_comm] +@[simp] lemma card_product : (s ×ˢ t).card = s.card * t.card := by simp [product] end product diff --git a/src/data/multiset/interval.lean b/src/data/multiset/interval.lean index af501ed5fd1bd..ce6951ea6a794 100644 --- a/src/data/multiset/interval.lean +++ b/src/data/multiset/interval.lean @@ -17,7 +17,7 @@ cardinality of its finite intervals. ## Implementation notes We implement the intervals via the intervals on `dfinsupp`, rather than via filtering -`multiset.powerset`; this is because `(multiset.repeat x n).powerset` has `2^n` entries not `n+1` +`multiset.powerset`; this is because `(multiset.replicate n x).powerset` has `2^n` entries not `n+1` entries as it contains duplicates. We do not go via `finsupp` as this would be noncomputable, and multisets are typically used computationally. diff --git a/src/data/multiset/nodup.lean b/src/data/multiset/nodup.lean index 1d1b639604903..d2eb04b019bdc 100644 --- a/src/data/multiset/nodup.lean +++ b/src/data/multiset/nodup.lean @@ -45,7 +45,7 @@ theorem not_nodup_pair : ∀ a : α, ¬ nodup (a ::ₘ a ::ₘ 0) := not_nodup_p theorem nodup_iff_le {s : multiset α} : nodup s ↔ ∀ a : α, ¬ a ::ₘ a ::ₘ 0 ≤ s := quot.induction_on s $ λ l, nodup_iff_sublist.trans $ forall_congr $ λ a, -not_congr (@repeat_le_coe _ a 2 _).symm + (@replicate_le_coe _ a 2 _).symm.not lemma nodup_iff_ne_cons_cons {s : multiset α} : s.nodup ↔ ∀ a t, s ≠ a ::ₘ a ::ₘ t := nodup_iff_le.trans diff --git a/src/data/nat/choose/multinomial.lean b/src/data/nat/choose/multinomial.lean index 6b7b4e402fd98..1670c3605d621 100644 --- a/src/data/nat/choose/multinomial.lean +++ b/src/data/nat/choose/multinomial.lean @@ -220,9 +220,9 @@ begin refine (fintype.sum_equiv (sym_insert_equiv ha) _ _ $ λ m, _).symm, rw [m.1.1.multinomial_filter_ne a], conv in (m.1.1.map _) { rw [← m.1.1.filter_add_not ((=) a), multiset.map_add] }, - simp_rw [multiset.noncomm_prod_add, m.1.1.filter_eq, multiset.map_repeat, m.1.2], - rw [multiset.noncomm_prod_eq_pow_card _ _ _ (λ _, multiset.eq_of_mem_repeat)], - rw [multiset.card_repeat, nat.cast_mul, mul_assoc, nat.cast_comm], + simp_rw [multiset.noncomm_prod_add, m.1.1.filter_eq, multiset.map_replicate, m.1.2], + rw [multiset.noncomm_prod_eq_pow_card _ _ _ (λ _, multiset.eq_of_mem_replicate)], + rw [multiset.card_replicate, nat.cast_mul, mul_assoc, nat.cast_comm], congr' 1, simp_rw [← mul_assoc, nat.cast_comm], refl, end diff --git a/src/data/nat/digits.lean b/src/data/nat/digits.lean index 8cb8a40447742..60d07ee5c227f 100644 --- a/src/data/nat/digits.lean +++ b/src/data/nat/digits.lean @@ -34,7 +34,7 @@ def digits_aux_0 : ℕ → list ℕ | (n+1) := [n+1] /-- (Impl.) An auxiliary definition for `digits`, to help get the desired definitional unfolding. -/ -def digits_aux_1 (n : ℕ) : list ℕ := list.repeat 1 n +def digits_aux_1 (n : ℕ) : list ℕ := list.replicate n 1 /-- (Impl.) An auxiliary definition for `digits`, to help get the desired definitional unfolding. -/ def digits_aux (b : ℕ) (h : 2 ≤ b) : ℕ → list ℕ @@ -61,7 +61,7 @@ In any base, we have `of_digits b L = L.foldr (λ x y, x + b * y) 0`. * For any `2 ≤ b`, we have `l < b` for any `l ∈ digits b n`, and the last digit is not zero. This uniquely specifies the behaviour of `digits b`. -* For `b = 1`, we define `digits 1 n = list.repeat 1 n`. +* For `b = 1`, we define `digits 1 n = list.replicate n 1`. * For `b = 0`, we define `digits 0 n = [n]`, except `digits 0 0 = []`. Note this differs from the existing `nat.to_digits` in core, which is used for printing numerals. @@ -83,7 +83,7 @@ theorem digits_zero_succ' : ∀ {n : ℕ} (w : 0 < n), digits 0 n = [n] | 0 h := absurd h dec_trivial | (n+1) _ := rfl -@[simp] lemma digits_one (n : ℕ) : digits 1 n = list.repeat 1 n := rfl +@[simp] lemma digits_one (n : ℕ) : digits 1 n = list.replicate n 1 := rfl @[simp] lemma digits_one_succ (n : ℕ) : digits 1 (n + 1) = 1 :: digits 1 n := rfl @@ -347,7 +347,7 @@ begin { cases hm rfl }, { simp } }, { cases m, { cases hm rfl }, - simp_rw [digits_one, list.last_repeat_succ 1 m], + simp_rw [digits_one, list.last_replicate_succ 1 m], norm_num }, revert hm, apply nat.strong_induction_on m, @@ -691,8 +691,8 @@ example : nat.digits 10 123 = [3,2,1] := by norm_num else if b = 1 then do ic ← mk_instance_cache `(ℕ), (_, pn0) ← norm_num.prove_pos ic en, - s ← simp_lemmas.add_simp simp_lemmas.mk `list.repeat, - (rhs, p2, _) ← simplify s [] `(list.repeat 1 %%en), + s ← simp_lemmas.add_simp simp_lemmas.mk `list.replicate, + (rhs, p2, _) ← simplify s [] `(list.replicate %%en 1), p ← mk_eq_trans `(nat.digits_one %%en) p2, return (rhs, p) else do diff --git a/src/data/nat/factorization/basic.lean b/src/data/nat/factorization/basic.lean index 950bdb874bac2..2e46c11e86c66 100644 --- a/src/data/nat/factorization/basic.lean +++ b/src/data/nat/factorization/basic.lean @@ -307,7 +307,7 @@ begin apply dvd_of_factors_subperm (pow_ne_zero _ hp.ne_zero), rw [hp.factors_pow, list.subperm_ext_iff], intros q hq, - simp [list.eq_of_mem_repeat hq], + simp [list.eq_of_mem_replicate hq], end lemma ord_compl_dvd (n p : ℕ) : ord_compl[p] n ∣ n := diff --git a/src/data/nat/factors.lean b/src/data/nat/factors.lean index b06dac80eaf84..20277b4343685 100644 --- a/src/data/nat/factors.lean +++ b/src/data/nat/factors.lean @@ -154,13 +154,13 @@ begin end lemma prime.factors_pow {p : ℕ} (hp : p.prime) (n : ℕ) : - (p ^ n).factors = list.repeat p n := + (p ^ n).factors = list.replicate n p := begin symmetry, - rw ← list.repeat_perm, - apply nat.factors_unique (list.prod_repeat p n), + rw ← list.replicate_perm, + apply nat.factors_unique (list.prod_replicate p n), intros q hq, - rwa eq_of_mem_repeat hq, + rwa eq_of_mem_replicate hq, end lemma eq_prime_pow_of_unique_prime_dvd {n p : ℕ} (hpos : n ≠ 0) @@ -168,8 +168,8 @@ lemma eq_prime_pow_of_unique_prime_dvd {n p : ℕ} (hpos : n ≠ 0) n = p ^ n.factors.length := begin set k := n.factors.length, - rw [←prod_factors hpos, ←prod_repeat p k, - eq_repeat_of_mem (λ d hd, h (prime_of_mem_factors hd) (dvd_of_mem_factors hd))], + rw [←prod_factors hpos, ←prod_replicate p k, + eq_replicate_of_mem (λ d hd, h (prime_of_mem_factors hd) (dvd_of_mem_factors hd))], end /-- For positive `a` and `b`, the prime factors of `a * b` are the union of those of `a` and `b` -/ diff --git a/src/data/pnat/factors.lean b/src/data/pnat/factors.lean index f49517117765b..10a72127cb164 100644 --- a/src/data/pnat/factors.lean +++ b/src/data/pnat/factors.lean @@ -345,14 +345,14 @@ theorem count_factor_multiset (m : ℕ+) (p : nat.primes) (k : ℕ) : (p : ℕ+) ^ k ∣ m ↔ k ≤ m.factor_multiset.count p := begin intros, - rw [multiset.le_count_iff_repeat_le], + rw [multiset.le_count_iff_replicate_le], rw [← factor_multiset_le_iff, factor_multiset_pow, factor_multiset_of_prime], congr' 2, - apply multiset.eq_repeat.mpr, + apply multiset.eq_replicate.mpr, split, { rw [multiset.card_nsmul, prime_multiset.card_of_prime, mul_one] }, { intros q h, rw [prime_multiset.of_prime, multiset.nsmul_singleton _ k] at h, - exact multiset.eq_of_mem_repeat h } + exact multiset.eq_of_mem_replicate h } end end pnat diff --git a/src/data/polynomial/ring_division.lean b/src/data/polynomial/ring_division.lean index bd861c1787f46..cb4c0ce487998 100644 --- a/src/data/polynomial/ring_division.lean +++ b/src/data/polynomial/ring_division.lean @@ -617,7 +617,7 @@ end (s.map (λ a, X - C a)).prod.nat_degree = s.card := begin rw [nat_degree_multiset_prod_of_monic, multiset.map_map], - { convert multiset.sum_repeat 1 _, + { convert multiset.sum_replicate 1 _, { convert multiset.map_const _ 1, ext, apply nat_degree_X_sub_C }, { simp } }, { intros f hf, obtain ⟨a, ha, rfl⟩ := multiset.mem_map.1 hf, exact monic_X_sub_C a }, end @@ -909,8 +909,8 @@ end lemma _root_.multiset.prod_X_sub_C_dvd_iff_le_roots {p : R[X]} (hp : p ≠ 0) (s : multiset R) : (s.map (λ a, X - C a)).prod ∣ p ↔ s ≤ p.roots := ⟨λ h, multiset.le_iff_count.2 $ λ r, begin - rw [count_roots, le_root_multiplicity_iff hp, ← multiset.prod_repeat, - ← multiset.map_repeat (λ a, X - C a), ← multiset.filter_eq], + rw [count_roots, le_root_multiplicity_iff hp, ← multiset.prod_replicate, + ← multiset.map_replicate (λ a, X - C a), ← multiset.filter_eq], exact (multiset.prod_dvd_prod_of_le $ multiset.map_le_map $ s.filter_le _).trans h, end, λ h, (multiset.prod_dvd_prod_of_le $ multiset.map_le_map h).trans p.prod_multiset_X_sub_C_dvd⟩ @@ -971,7 +971,8 @@ end lemma count_map_roots [is_domain A] {p : A[X]} {f : A →+* B} (hmap : map f p ≠ 0) (b : B) : (p.roots.map f).count b ≤ root_multiplicity b (p.map f) := begin - rw [le_root_multiplicity_iff hmap, ← multiset.prod_repeat, ← multiset.map_repeat (λ a, X - C a)], + rw [le_root_multiplicity_iff hmap, ← multiset.prod_replicate, + ← multiset.map_replicate (λ a, X - C a)], rw ← multiset.filter_eq, refine (multiset.prod_dvd_prod_of_le $ multiset.map_le_map $ multiset.filter_le _ _).trans _, convert polynomial.map_dvd _ p.prod_multiset_X_sub_C_dvd, diff --git a/src/data/polynomial/splits.lean b/src/data/polynomial/splits.lean index a28f5272e3995..6d98f5f312386 100644 --- a/src/data/polynomial/splits.lean +++ b/src/data/polynomial/splits.lean @@ -403,7 +403,7 @@ begin simp_rw [function.comp_app, eval_sub, eval_X, zero_sub, eval_C], conv_lhs { congr, congr, funext, rw [neg_eq_neg_one_mul] }, - rw [multiset.prod_map_mul, multiset.map_const, multiset.prod_repeat, multiset.map_id', + rw [multiset.prod_map_mul, multiset.map_const, multiset.prod_replicate, multiset.map_id', splits_iff_card_roots.1 hP] end diff --git a/src/data/qpf/multivariate/constructions/cofix.lean b/src/data/qpf/multivariate/constructions/cofix.lean index 3023b1e0fc374..ba67e60e0df8b 100644 --- a/src/data/qpf/multivariate/constructions/cofix.lean +++ b/src/data/qpf/multivariate/constructions/cofix.lean @@ -423,7 +423,7 @@ do e ← to_expr e, R ← pose `R none ex, refine ``(cofix.bisim₂ %%R _ _ _ ⟨_,rfl,rfl⟩), let f (a b : name) : name := if a = `_ then b else a, - let ids := (ids ++ list.repeat `_ 5).zip_with f [`a,`b,`x,`Ha,`Hb], + let ids := (ids ++ list.replicate 5 `_).zip_with f [`a,`b,`x,`Ha,`Hb], (ids₀,w::ids₁) ← pure $ list.split_at 2 ids, intro_lst ids₀, h ← intro1, diff --git a/src/data/set/pointwise/list_of_fn.lean b/src/data/set/pointwise/list_of_fn.lean index 3356bf4453b14..3019064cf9ff5 100644 --- a/src/data/set/pointwise/list_of_fn.lean +++ b/src/data/set/pointwise/list_of_fn.lean @@ -44,6 +44,6 @@ end @[to_additive] lemma mem_pow {a : α} {n : ℕ} : a ∈ s ^ n ↔ ∃ f : fin n → s, (list.of_fn (λ i, (f i : α))).prod = a := -by rw [←mem_prod_list_of_fn, list.of_fn_const, list.prod_repeat] +by rw [←mem_prod_list_of_fn, list.of_fn_const, list.prod_replicate] end set diff --git a/src/data/sym/basic.lean b/src/data/sym/basic.lean index 556b140d654e5..f1a70a1fe9a65 100644 --- a/src/data/sym/basic.lean +++ b/src/data/sym/basic.lean @@ -185,20 +185,19 @@ subtype.ext $ multiset.card_eq_zero.1 s.2 instance unique_zero : unique (sym α 0) := ⟨⟨nil⟩, eq_nil_of_card_zero⟩ -/-- `repeat a n` is the sym containing only `a` with multiplicity `n`. -/ -def repeat (a : α) (n : ℕ) : sym α n := ⟨multiset.repeat a n, multiset.card_repeat _ _⟩ +/-- `replicate n a` is the sym containing only `a` with multiplicity `n`. -/ +def replicate (n : ℕ) (a : α) : sym α n := ⟨multiset.replicate n a, multiset.card_replicate _ _⟩ -lemma repeat_succ {a : α} {n : ℕ} : repeat a n.succ = a ::ₛ repeat a n := rfl +lemma replicate_succ {a : α} {n : ℕ} : replicate n.succ a = a ::ₛ replicate n a := rfl -lemma coe_repeat : (repeat a n : multiset α) = multiset.repeat a n := rfl +lemma coe_replicate : (replicate n a : multiset α) = multiset.replicate n a := rfl -@[simp] lemma mem_repeat : b ∈ repeat a n ↔ n ≠ 0 ∧ b = a := multiset.mem_repeat +@[simp] lemma mem_replicate : b ∈ replicate n a ↔ n ≠ 0 ∧ b = a := multiset.mem_replicate -lemma eq_repeat_iff : s = repeat a n ↔ ∀ b ∈ s, b = a := +lemma eq_replicate_iff : s = replicate n a ↔ ∀ b ∈ s, b = a := begin - rw [subtype.ext_iff, coe_repeat], - convert multiset.eq_repeat', - exact s.2.symm, + rw [subtype.ext_iff, coe_replicate, multiset.eq_replicate], + exact and_iff_right s.2 end lemma exists_mem (s : sym α n.succ) : ∃ a, a ∈ s := @@ -211,11 +210,12 @@ begin exact ⟨a, s.erase a ha, (cons_erase ha).symm⟩, end -lemma eq_repeat {a : α} {n : ℕ} {s : sym α n} : s = repeat a n ↔ ∀ b ∈ s, b = a := -subtype.ext_iff.trans $ multiset.eq_repeat.trans $ and_iff_right s.prop +lemma eq_replicate {a : α} {n : ℕ} {s : sym α n} : s = replicate n a ↔ ∀ b ∈ s, b = a := +subtype.ext_iff.trans $ multiset.eq_replicate.trans $ and_iff_right s.prop -lemma eq_repeat_of_subsingleton [subsingleton α] (a : α) {n : ℕ} (s : sym α n) : s = repeat a n := -eq_repeat.2 $ λ b hb, subsingleton.elim _ _ +lemma eq_replicate_of_subsingleton [subsingleton α] (a : α) {n : ℕ} (s : sym α n) : + s = replicate n a := +eq_replicate.2 $ λ b hb, subsingleton.elim _ _ instance [subsingleton α] (n : ℕ) : subsingleton (sym α n) := ⟨begin @@ -223,28 +223,29 @@ instance [subsingleton α] (n : ℕ) : subsingleton (sym α n) := { simp, }, { intros s s', obtain ⟨b, -⟩ := exists_mem s, - rw [eq_repeat_of_subsingleton b s', eq_repeat_of_subsingleton b s], }, + rw [eq_replicate_of_subsingleton b s', eq_replicate_of_subsingleton b s], }, end⟩ instance inhabited_sym [inhabited α] (n : ℕ) : inhabited (sym α n) := -⟨repeat default n⟩ +⟨replicate n default⟩ instance inhabited_sym' [inhabited α] (n : ℕ) : inhabited (sym' α n) := -⟨quotient.mk' (vector.repeat default n)⟩ +⟨quotient.mk' (vector.replicate n default)⟩ instance (n : ℕ) [is_empty α] : is_empty (sym α n.succ) := ⟨λ s, by { obtain ⟨a, -⟩ := exists_mem s, exact is_empty_elim a }⟩ instance (n : ℕ) [unique α] : unique (sym α n) := unique.mk' _ -lemma repeat_left_inj {a b : α} {n : ℕ} (h : n ≠ 0) : repeat a n = repeat b n ↔ a = b := -subtype.ext_iff.trans (multiset.repeat_left_inj h) +lemma replicate_right_inj {a b : α} {n : ℕ} (h : n ≠ 0) : replicate n a = replicate n b ↔ a = b := +subtype.ext_iff.trans (multiset.replicate_right_inj h) -lemma repeat_left_injective {n : ℕ} (h : n ≠ 0) : function.injective (λ x : α, repeat x n) := -λ a b, (repeat_left_inj h).1 +lemma replicate_right_injective {n : ℕ} (h : n ≠ 0) : + function.injective (replicate n : α → sym α n) := +λ a b, (replicate_right_inj h).1 instance (n : ℕ) [nontrivial α] : nontrivial (sym α (n + 1)) := -(repeat_left_injective n.succ_ne_zero).nontrivial +(replicate_right_injective n.succ_ne_zero).nontrivial /-- A function `α → β` induces a function `sym α n → sym β n` by applying it to every element of the underlying `n`-tuple. -/ @@ -355,16 +356,17 @@ by { ext, simp [append, add_comm], } lemma mem_append_iff {s' : sym α m} : a ∈ s.append s' ↔ a ∈ s ∨ a ∈ s' := multiset.mem_add /-- Fill a term `m : sym α (n - i)` with `i` copies of `a` to obtain a term of `sym α n`. -This is a convenience wrapper for `m.append (repeat a i)` that adjusts the term using `sym.cast`. -/ +This is a convenience wrapper for `m.append (replicate i a)` that adjusts the term using +`sym.cast`. -/ def fill (a : α) (i : fin (n + 1)) (m : sym α (n - i)) : sym α n := -sym.cast (nat.sub_add_cancel i.is_le) (m.append (repeat a i)) +sym.cast (nat.sub_add_cancel i.is_le) (m.append (replicate i a)) lemma coe_fill {a : α} {i : fin (n + 1)} {m : sym α (n - i)} : - (fill a i m : multiset α) = m + repeat a i := rfl + (fill a i m : multiset α) = m + replicate i a := rfl lemma mem_fill_iff {a b : α} {i : fin (n + 1)} {s : sym α (n - i)} : a ∈ sym.fill b i s ↔ ((i : ℕ) ≠ 0 ∧ a = b) ∨ a ∈ s := -by rw [fill, mem_cast, mem_append_iff, or_comm, mem_repeat] +by rw [fill, mem_cast, mem_append_iff, or_comm, mem_replicate] open multiset @@ -386,7 +388,7 @@ lemma fill_filter_ne [decidable_eq α] (a : α) (m : sym α n) : (m.filter_ne a).2.fill a (m.filter_ne a).1 = m := subtype.ext begin dsimp only [coe_fill, filter_ne, subtype.coe_mk, fin.coe_mk], - ext b, rw [count_add, count_filter, sym.coe_repeat, count_repeat], + ext b, rw [count_add, count_filter, sym.coe_replicate, count_replicate], obtain rfl | h := eq_or_ne a b, { rw [if_pos rfl, if_neg (not_not.2 rfl), zero_add], refl }, { rw [if_pos h, if_neg h.symm, add_zero], refl }, @@ -397,7 +399,7 @@ lemma filter_ne_fill [decidable_eq α] (a : α) (m : Σ i : fin (n + 1), sym α sigma_sub_ext begin dsimp only [filter_ne, subtype.coe_mk, subtype.val_eq_coe, coe_fill], rw [filter_add, filter_eq_self.2, add_right_eq_self, eq_zero_iff_forall_not_mem], - { intros b hb, rw [mem_filter, sym.mem_coe, mem_repeat] at hb, exact hb.2 hb.1.2.symm }, + { intros b hb, rw [mem_filter, sym.mem_coe, mem_replicate] at hb, exact hb.2 hb.1.2.symm }, { exact λ b hb, (hb.ne_of_not_mem h).symm }, end diff --git a/src/data/vector/basic.lean b/src/data/vector/basic.lean index d958b566e2f0c..8aa74ead6d4e2 100644 --- a/src/data/vector/basic.lean +++ b/src/data/vector/basic.lean @@ -35,6 +35,10 @@ subtype.val_injective | ⟨v, hv⟩ ⟨w, hw⟩ h := subtype.eq (list.ext_le (by rw [hv, hw]) (λ m hm hn, h ⟨m, hv ▸ hm⟩)) +/-- A vector with `n` elements `a`. -/ +def replicate (n : ℕ) (a : α) : vector α n := +⟨list.replicate n a, list.length_replicate n a⟩ + /-- The empty `vector` is a `subsingleton`. -/ instance zero_subsingleton : subsingleton (vector α 0) := ⟨λ _ _, vector.ext (λ m, fin.elim0 m)⟩ @@ -96,9 +100,9 @@ theorem nth_eq_nth_le : ∀ (v : vector α n) (i), | ⟨l, h⟩ i := rfl @[simp] -lemma nth_repeat (a : α) (i : fin n) : - (vector.repeat a n).nth i = a := -by apply list.nth_le_repeat +lemma nth_replicate (a : α) (i : fin n) : + (vector.replicate n a).nth i = a := +list.nth_le_replicate _ _ @[simp] lemma nth_map {β : Type*} (v : vector α n) (f : α → β) (i : fin n) : (v.map f).nth i = f (v.nth i) := diff --git a/src/field_theory/abel_ruffini.lean b/src/field_theory/abel_ruffini.lean index c637f0cce91a7..d5d694479963d 100644 --- a/src/field_theory/abel_ruffini.lean +++ b/src/field_theory/abel_ruffini.lean @@ -184,7 +184,7 @@ begin change (X - C c).comp (C b * X) = C b * (X - C (c / b)), rw [sub_comp, X_comp, C_comp, mul_sub, ←C_mul, mul_div_cancel' c hb'] }, rw [key1, hs, multiset_prod_comp, multiset.map_map, key2, multiset.prod_map_mul, - multiset.map_const, multiset.prod_repeat, hs', ←C_pow, hb, ←mul_assoc, C_mul_C, one_mul], + multiset.map_const, multiset.prod_replicate, hs', ←C_pow, hb, ←mul_assoc, C_mul_C, one_mul], all_goals { exact field.to_nontrivial F }, end diff --git a/src/group_theory/p_group.lean b/src/group_theory/p_group.lean index 63e4e9ae7a54c..c07f252b002d1 100644 --- a/src/group_theory/p_group.lean +++ b/src/group_theory/p_group.lean @@ -52,7 +52,7 @@ begin refine ⟨λ h, _, λ ⟨n, hn⟩, of_card hn⟩, suffices : ∀ q ∈ nat.factors (card G), q = p, { use (card G).factors.length, - rw [←list.prod_repeat, ←list.eq_repeat_of_mem this, nat.prod_factors hG] }, + rw [←list.prod_replicate, ←list.eq_replicate_of_mem this, nat.prod_factors hG] }, intros q hq, obtain ⟨hq1, hq2⟩ := (nat.mem_factors hG).mp hq, haveI : fact q.prime := ⟨hq1⟩, diff --git a/src/group_theory/perm/cycle/type.lean b/src/group_theory/perm/cycle/type.lean index 1aac2fe826279..cb01ef0cec419 100644 --- a/src/group_theory/perm/cycle/type.lean +++ b/src/group_theory/perm/cycle/type.lean @@ -206,9 +206,9 @@ lemma two_dvd_card_support {σ : perm α} (hσ : σ ^ 2 = 1) : 2 ∣ σ.support. (dvd_of_mem_cycle_type hn).trans $ order_of_dvd_of_pow_eq_one hσ) (two_le_of_mem_cycle_type hn))) lemma cycle_type_prime_order {σ : perm α} (hσ : (order_of σ).prime) : - ∃ n : ℕ, σ.cycle_type = repeat (order_of σ) (n + 1) := + ∃ n : ℕ, σ.cycle_type = replicate (n + 1) (order_of σ) := begin - rw eq_repeat_of_mem (λ n hn, or_iff_not_imp_left.mp + rw eq_replicate_of_mem (λ n hn, or_iff_not_imp_left.mp (hσ.eq_one_or_self_of_dvd n (dvd_of_mem_cycle_type hn)) (one_lt_of_mem_cycle_type hn).ne'), use σ.cycle_type.card - 1, rw tsub_add_cancel_of_le, @@ -222,9 +222,9 @@ lemma is_cycle_of_prime_order {σ : perm α} (h1 : (order_of σ).prime) (h2 : σ.support.card < 2 * (order_of σ)) : σ.is_cycle := begin obtain ⟨n, hn⟩ := cycle_type_prime_order h1, - rw [←σ.sum_cycle_type, hn, multiset.sum_repeat, nsmul_eq_mul, nat.cast_id, mul_lt_mul_right + rw [←σ.sum_cycle_type, hn, multiset.sum_replicate, nsmul_eq_mul, nat.cast_id, mul_lt_mul_right (order_of_pos σ), nat.succ_lt_succ_iff, nat.lt_succ_iff, le_zero_iff] at h2, - rw [←card_cycle_type_eq_one, hn, card_repeat, h2], + rw [←card_cycle_type_eq_one, hn, card_replicate, h2], end lemma cycle_type_le_of_mem_cycle_factors_finset {f g : perm α} @@ -481,12 +481,13 @@ begin have hσ : ∀ k v, (σ ^ k) v = f k v := λ k v, nat.rec (hf1 v).symm (λ k hk, eq.trans (by exact congr_arg σ hk) (hf2 k 1 v)) k, replace hσ : σ ^ (p ^ 1) = 1 := perm.ext (λ v, by rw [pow_one, hσ, hf3, one_apply]), - let v₀ : vectors_prod_eq_one G p := ⟨vector.repeat 1 p, (list.prod_repeat 1 p).trans (one_pow p)⟩, - have hv₀ : σ v₀ = v₀ := subtype.ext (subtype.ext (list.rotate_repeat (1 : G) p 1)), + let v₀ : vectors_prod_eq_one G p := + ⟨vector.replicate p 1, (list.prod_replicate 1 p).trans (one_pow p)⟩, + have hv₀ : σ v₀ = v₀ := subtype.ext (subtype.ext (list.rotate_replicate (1 : G) p 1)), obtain ⟨v, hv1, hv2⟩ := exists_fixed_point_of_prime' Scard hσ hv₀, refine exists_imp_exists (λ g hg, order_of_eq_prime _ (λ hg', hv2 _)) - (list.rotate_one_eq_self_iff_eq_repeat.mp (subtype.ext_iff.mp (subtype.ext_iff.mp hv1))), - { rw [←list.prod_repeat, ←v.1.2, ←hg, (show v.val.val.prod = 1, from v.2)] }, + (list.rotate_one_eq_self_iff_eq_replicate.mp (subtype.ext_iff.mp (subtype.ext_iff.mp hv1))), + { rw [←list.prod_replicate, ←v.1.2, ←hg, (show v.val.val.prod = 1, from v.2)] }, { rw [subtype.ext_iff_val, subtype.ext_iff_val, hg, hg', v.1.2], refl }, end @@ -523,25 +524,25 @@ variables [decidable_eq α] /-- The partition corresponding to a permutation -/ def partition (σ : perm α) : (fintype.card α).partition := -{ parts := σ.cycle_type + repeat 1 (fintype.card α - σ.support.card), +{ parts := σ.cycle_type + replicate (fintype.card α - σ.support.card) 1, parts_pos := λ n hn, begin cases mem_add.mp hn with hn hn, { exact zero_lt_one.trans (one_lt_of_mem_cycle_type hn) }, - { exact lt_of_lt_of_le zero_lt_one (ge_of_eq (multiset.eq_of_mem_repeat hn)) }, + { exact lt_of_lt_of_le zero_lt_one (ge_of_eq (multiset.eq_of_mem_replicate hn)) }, end, - parts_sum := by rw [sum_add, sum_cycle_type, multiset.sum_repeat, nsmul_eq_mul, + parts_sum := by rw [sum_add, sum_cycle_type, multiset.sum_replicate, nsmul_eq_mul, nat.cast_id, mul_one, add_tsub_cancel_of_le σ.support.card_le_univ] } lemma parts_partition {σ : perm α} : - σ.partition.parts = σ.cycle_type + repeat 1 (fintype.card α - σ.support.card) := rfl + σ.partition.parts = σ.cycle_type + replicate (fintype.card α - σ.support.card) 1 := rfl lemma filter_parts_partition_eq_cycle_type {σ : perm α} : (partition σ).parts.filter (λ n, 2 ≤ n) = σ.cycle_type := begin rw [parts_partition, filter_add, multiset.filter_eq_self.2 (λ _, two_le_of_mem_cycle_type), multiset.filter_eq_nil.2 (λ a h, _), add_zero], - rw multiset.eq_of_mem_repeat h, + rw multiset.eq_of_mem_replicate h, dec_trivial end diff --git a/src/group_theory/perm/sign.lean b/src/group_theory/perm/sign.lean index 8ac87f929df47..77d8db5346ed2 100644 --- a/src/group_theory/perm/sign.lean +++ b/src/group_theory/perm/sign.lean @@ -518,11 +518,11 @@ sign_symm_trans_trans f e.symm lemma sign_prod_list_swap {l : list (perm α)} (hl : ∀ g ∈ l, is_swap g) : sign l.prod = (-1) ^ l.length := -have h₁ : l.map sign = list.repeat (-1) l.length := - list.eq_repeat.2 ⟨by simp, λ u hu, +have h₁ : l.map sign = list.replicate l.length (-1) := + list.eq_replicate.2 ⟨by simp, λ u hu, let ⟨g, hg⟩ := list.mem_map.1 hu in hg.2 ▸ (hl _ hg.1).sign_eq⟩, -by rw [← list.prod_repeat, ← h₁, list.prod_hom _ (@sign α _ _)] +by rw [← list.prod_replicate, ← h₁, list.prod_hom _ (@sign α _ _)] variable (α) @@ -545,15 +545,15 @@ have ∀ {f}, is_swap f → s f = -1 := have ∀ a ∈ l.map s, a = (1 : ℤˣ) := λ a ha, let ⟨g, hg⟩ := list.mem_map.1 ha in hg.2 ▸ this _ (hl.2 _ hg.1), have s l.prod = 1, - by rw [← l.prod_hom s, list.eq_repeat'.2 this, list.prod_repeat, one_pow], + by rw [← l.prod_hom s, list.eq_replicate_length.2 this, list.prod_replicate, one_pow], by { rw [hl.1, hg] at this, exact absurd this dec_trivial }), monoid_hom.ext $ λ f, let ⟨l, hl₁, hl₂⟩ := (trunc_swap_factors f).out in have hsl : ∀ a ∈ l.map s, a = (-1 : ℤˣ) := λ a ha, let ⟨g, hg⟩ := list.mem_map.1 ha in hg.2 ▸ this (hl₂ _ hg.1), -by rw [← hl₁, ← l.prod_hom s, list.eq_repeat'.2 hsl, list.length_map, - list.prod_repeat, sign_prod_list_swap hl₂] +by rw [← hl₁, ← l.prod_hom s, list.eq_replicate_length.2 hsl, list.length_map, + list.prod_replicate, sign_prod_list_swap hl₂] lemma sign_subtype_perm (f : perm α) {p : α → Prop} [decidable_pred p] (h₁ : ∀ x, p x ↔ p (f x)) (h₂ : ∀ x, f x ≠ x → p x) : sign (subtype_perm f h₁) = sign f := diff --git a/src/group_theory/specific_groups/alternating.lean b/src/group_theory/specific_groups/alternating.lean index 70da8978115a6..5fddc8b655f95 100644 --- a/src/group_theory/specific_groups/alternating.lean +++ b/src/group_theory/specific_groups/alternating.lean @@ -260,10 +260,9 @@ lemma is_conj_swap_mul_swap_of_cycle_type_two {g : perm (fin 5)} is_conj (swap 0 4 * swap 1 3) g := begin have h := g.support.card_le_univ, - rw [← sum_cycle_type, multiset.eq_repeat_of_mem h2, multiset.sum_repeat, smul_eq_mul] at h, - rw [← multiset.eq_repeat'] at h2, - have h56 : 5 ≤ 3 * 2 := nat.le_succ 5, - have h := le_of_mul_le_mul_right (le_trans h h56) dec_trivial, + rw [← multiset.eq_replicate_card] at h2, + rw [← sum_cycle_type, h2, multiset.sum_replicate, smul_eq_mul] at h, + have h : g.cycle_type.card ≤ 3 := le_of_mul_le_mul_right (le_trans h dec_trivial) dec_trivial, rw [mem_alternating_group, sign_of_cycle_type, h2] at ha, norm_num at ha, rw [pow_add, pow_mul, int.units_pow_two,one_mul, diff --git a/src/logic/equiv/list.lean b/src/logic/equiv/list.lean index f0a46309606e7..63cda4d0f5634 100644 --- a/src/logic/equiv/list.lean +++ b/src/logic/equiv/list.lean @@ -344,9 +344,9 @@ namespace equiv /-- The type lists on unit is canonically equivalent to the natural numbers. -/ def list_unit_equiv : list unit ≃ ℕ := { to_fun := list.length, - inv_fun := list.repeat (), + inv_fun := λ n, list.replicate n (), left_inv := λ u, list.length_injective (by simp), - right_inv := λ n, list.length_repeat () n } + right_inv := λ n, list.length_replicate n () } /-- `list ℕ` is equivalent to `ℕ`. -/ def list_nat_equiv_nat : list ℕ ≃ ℕ := denumerable.eqv _ diff --git a/src/number_theory/arithmetic_function.lean b/src/number_theory/arithmetic_function.lean index 2d05ecfe97d50..332b45f5f2cd3 100644 --- a/src/number_theory/arithmetic_function.lean +++ b/src/number_theory/arithmetic_function.lean @@ -732,7 +732,7 @@ end card_factors_eq_one_iff_prime.2 hp @[simp] lemma card_factors_apply_prime_pow {p k : ℕ} (hp : p.prime) : Ω (p ^ k) = k := -by rw [card_factors_apply, hp.factors_pow, list.length_repeat] +by rw [card_factors_apply, hp.factors_pow, list.length_replicate] /-- `ω n` is the number of distinct prime factors of `n`. -/ def card_distinct_factors : arithmetic_function ℕ := @@ -761,7 +761,7 @@ end @[simp] lemma card_distinct_factors_apply_prime_pow {p k : ℕ} (hp : p.prime) (hk : k ≠ 0) : ω (p ^ k) = 1 := -by rw [card_distinct_factors_apply, hp.factors_pow, list.repeat_dedup hk, list.length_singleton] +by rw [card_distinct_factors_apply, hp.factors_pow, list.replicate_dedup hk, list.length_singleton] @[simp] lemma card_distinct_factors_apply_prime {p : ℕ} (hp : p.prime) : ω p = 1 := by rw [←pow_one p, card_distinct_factors_apply_prime_pow hp one_ne_zero] diff --git a/src/number_theory/ramification_inertia.lean b/src/number_theory/ramification_inertia.lean index f957423607628..4c77e129f0147 100644 --- a/src/number_theory/ramification_inertia.lean +++ b/src/number_theory/ramification_inertia.lean @@ -142,7 +142,7 @@ begin refine ramification_idx_spec (ideal.le_of_dvd _) (mt ideal.dvd_iff_le.mpr _); rw [dvd_iff_normalized_factors_le_normalized_factors (pow_ne_zero _ hP0) hp0, normalized_factors_pow, normalized_factors_irreducible hPirr, normalize_eq, - multiset.nsmul_singleton, ← multiset.le_count_iff_repeat_le], + multiset.nsmul_singleton, ← multiset.le_count_iff_replicate_le], { exact (nat.lt_succ_self _).not_le }, end @@ -581,7 +581,7 @@ begin letI := classical.dec_eq (ideal S), rw [sup_eq_prod_inf_factors _ (pow_ne_zero _ hP0), normalized_factors_pow, normalized_factors_irreducible ((ideal.prime_iff_is_prime hP0).mpr hP).irreducible, - normalize_eq, multiset.nsmul_singleton, multiset.inter_repeat, multiset.prod_repeat], + normalize_eq, multiset.nsmul_singleton, multiset.inter_replicate, multiset.prod_replicate], rw [← submodule.span_singleton_le_iff_mem, ideal.submodule_span_eq] at a_mem a_not_mem, rwa [ideal.count_normalized_factors_eq a_mem a_not_mem, min_eq_left i.le_succ], { intro ha, diff --git a/src/ring_theory/chain_of_divisors.lean b/src/ring_theory/chain_of_divisors.lean index a64688573412a..8a729133e178c 100644 --- a/src/ring_theory/chain_of_divisors.lean +++ b/src/ring_theory/chain_of_divisors.lean @@ -146,14 +146,14 @@ lemma element_of_chain_eq_pow_second_of_chain {q r : associates M} {n : ℕ} (hn begin classical, let i := (normalized_factors r).card, - have hi : normalized_factors r = multiset.repeat (c 1) i, - { apply multiset.eq_repeat_of_mem, + have hi : normalized_factors r = multiset.replicate i (c 1), + { apply multiset.eq_replicate_of_mem, intros b hb, refine eq_second_of_chain_of_prime_dvd hn h₁ (λ r', h₂) (prime_of_normalized_factor b hb) hr (dvd_of_mem_normalized_factors hb) }, have H : r = (c 1)^i, { have := unique_factorization_monoid.normalized_factors_prod (ne_zero_of_dvd_ne_zero hq hr), - rw [associated_iff_eq, hi, multiset.prod_repeat] at this, + rw [associated_iff_eq, hi, multiset.prod_replicate] at this, rw this }, refine ⟨⟨i, _⟩, H⟩, diff --git a/src/ring_theory/dedekind_domain/ideal.lean b/src/ring_theory/dedekind_domain/ideal.lean index 9adff7d3d09cf..5765fa9516407 100644 --- a/src/ring_theory/dedekind_domain/ideal.lean +++ b/src/ring_theory/dedekind_domain/ideal.lean @@ -736,7 +736,7 @@ begin have := pow_ne_zero (i + 1) hP, rw [← ideal.dvd_not_unit_iff_lt, dvd_not_unit_iff_normalized_factors_lt_normalized_factors, normalized_factors_pow, normalized_factors_irreducible P_prime'.irreducible, - multiset.nsmul_singleton, multiset.lt_repeat_succ] + multiset.nsmul_singleton, multiset.lt_replicate_succ] at hlt, rw [← ideal.dvd_iff_le, dvd_iff_normalized_factors_le_normalized_factors, normalized_factors_pow, normalized_factors_irreducible P_prime'.irreducible, multiset.nsmul_singleton], @@ -907,8 +907,8 @@ end lemma irreducible_pow_sup (hI : I ≠ ⊥) (hJ : irreducible J) (n : ℕ) : J^n ⊔ I = J^(min ((normalized_factors I).count J) n) := -by rw [sup_eq_prod_inf_factors (pow_ne_zero n hJ.ne_zero) hI, ← inf_eq_inter, - normalized_factors_of_irreducible_pow hJ, normalize_eq J, repeat_inf, prod_repeat] +by rw [sup_eq_prod_inf_factors (pow_ne_zero n hJ.ne_zero) hI, min_comm, + normalized_factors_of_irreducible_pow hJ, normalize_eq J, replicate_inter, prod_replicate] lemma irreducible_pow_sup_of_le (hJ : irreducible J) (n : ℕ) (hn : ↑n ≤ multiplicity J I) : J^n ⊔ I = J^n := diff --git a/src/ring_theory/discrete_valuation_ring.lean b/src/ring_theory/discrete_valuation_ring.lean index b97fe3a6ba47f..9450069c0d4b8 100644 --- a/src/ring_theory/discrete_valuation_ring.lean +++ b/src/ring_theory/discrete_valuation_ring.lean @@ -192,10 +192,10 @@ let p := classical.some hR in let spec := classical.some_spec hR in unique_factorization_monoid.of_exists_prime_factors $ λ x hx, begin - use multiset.repeat p (classical.some (spec.2 hx)), + use multiset.replicate (classical.some (spec.2 hx)) p, split, { intros q hq, - have hpq := multiset.eq_of_mem_repeat hq, + have hpq := multiset.eq_of_mem_replicate hq, rw hpq, refine ⟨spec.1.ne_zero, spec.1.not_unit, _⟩, intros a b h, @@ -210,7 +210,7 @@ begin obtain ⟨m, rfl⟩ := nat.exists_eq_succ_of_ne_zero hm, rw pow_succ, apply dvd_mul_of_dvd_left dvd_rfl _ }, - { rw [multiset.prod_repeat], exact (classical.some_spec (spec.2 hx)), } + { rw [multiset.prod_replicate], exact (classical.some_spec (spec.2 hx)), } end omit hR @@ -227,10 +227,10 @@ begin refine ⟨fx.card, _⟩, have H := hfx.2, rw ← associates.mk_eq_mk_iff_associated at H ⊢, - rw [← H, ← associates.prod_mk, associates.mk_pow, ← multiset.prod_repeat], + rw [← H, ← associates.prod_mk, associates.mk_pow, ← multiset.prod_replicate], congr' 1, symmetry, - rw multiset.eq_repeat, + rw multiset.eq_replicate, simp only [true_and, and_imp, multiset.card_map, eq_self_iff_true, multiset.mem_map, exists_imp_distrib], rintros _ q hq rfl, @@ -329,9 +329,9 @@ begin unfreezingI { use fx.card }, have H := hfx.2, rw ← associates.mk_eq_mk_iff_associated at H ⊢, - rw [← H, ← associates.prod_mk, associates.mk_pow, ← multiset.prod_repeat], + rw [← H, ← associates.prod_mk, associates.mk_pow, ← multiset.prod_replicate], congr' 1, - rw multiset.eq_repeat, + rw multiset.eq_replicate, simp only [true_and, and_imp, multiset.card_map, eq_self_iff_true, multiset.mem_map, exists_imp_distrib], rintros _ _ _ rfl, @@ -368,16 +368,16 @@ lemma unit_mul_pow_congr_pow {p q : R} (hp : irreducible p) (hq : irreducible q) (u v : Rˣ) (m n : ℕ) (h : ↑u * p ^ m = v * q ^ n) : m = n := begin - have key : associated (multiset.repeat p m).prod (multiset.repeat q n).prod, - { rw [multiset.prod_repeat, multiset.prod_repeat, associated], + have key : associated (multiset.replicate m p).prod (multiset.replicate n q).prod, + { rw [multiset.prod_replicate, multiset.prod_replicate, associated], refine ⟨u * v⁻¹, _⟩, simp only [units.coe_mul], rw [mul_left_comm, ← mul_assoc, h, mul_right_comm, units.mul_inv, one_mul], }, have := multiset.card_eq_card_of_rel (unique_factorization_monoid.factors_unique _ _ key), - { simpa only [multiset.card_repeat] }, + { simpa only [multiset.card_replicate] }, all_goals { intros x hx, - unfreezingI { obtain rfl := multiset.eq_of_mem_repeat hx, assumption } }, + unfreezingI { obtain rfl := multiset.eq_of_mem_replicate hx, assumption } }, end lemma unit_mul_pow_congr_unit {ϖ : R} (hirr : irreducible ϖ) (u v : Rˣ) (m n : ℕ) diff --git a/src/ring_theory/polynomial/vieta.lean b/src/ring_theory/polynomial/vieta.lean index 529f82e0f12cc..2ae9290898d00 100644 --- a/src/ring_theory/polynomial/vieta.lean +++ b/src/ring_theory/polynomial/vieta.lean @@ -46,7 +46,7 @@ begin rw [esymm, ←sum_hom', ←sum_map_mul_right, map_congr (eq.refl _)], intros _ ht, rw mem_powerset_len at ht, - simp [ht, map_const, prod_repeat, prod_hom', map_id', card_sub], + simp [ht, map_const, prod_replicate, prod_hom', map_id', card_sub], end /-- Vieta's formula for the coefficients of the product of linear terms `X + λ` where `λ` runs @@ -88,7 +88,7 @@ begin rw [esymm, esymm, ←multiset.sum_map_mul_left, multiset.powerset_len_map, multiset.map_map, map_congr (eq.refl _)], intros x hx, - rw [(by { exact (mem_powerset_len.mp hx).right.symm }), ←prod_repeat, ←multiset.map_const], + rw [(by { exact (mem_powerset_len.mp hx).right.symm }), ←prod_replicate, ←multiset.map_const], nth_rewrite 2 ←map_id' x, rw [←prod_map_mul, map_congr (eq.refl _)], exact λ z _, neg_one_mul z, diff --git a/src/ring_theory/unique_factorization_domain.lean b/src/ring_theory/unique_factorization_domain.lean index 54c05dae5aa65..96f1548a897b8 100644 --- a/src/ring_theory/unique_factorization_domain.lean +++ b/src/ring_theory/unique_factorization_domain.lean @@ -626,7 +626,7 @@ begin end theorem _root_.irreducible.normalized_factors_pow {p : α} (hp : irreducible p) (k : ℕ) : - normalized_factors (p ^ k) = multiset.repeat (normalize p) k := + normalized_factors (p ^ k) = multiset.replicate k (normalize p) := by rw [normalized_factors_pow, normalized_factors_irreducible hp, multiset.nsmul_singleton] theorem normalized_factors_prod_eq (s : multiset α) (hs : ∀ a ∈ s, irreducible a) : @@ -666,7 +666,7 @@ begin end theorem normalized_factors_of_irreducible_pow {p : α} (hp : irreducible p) (k : ℕ) : - normalized_factors (p ^ k) = multiset.repeat (normalize p) k := + normalized_factors (p ^ k) = multiset.replicate k (normalize p) := by rw [normalized_factors_pow, normalized_factors_irreducible hp, multiset.nsmul_singleton] lemma zero_not_mem_normalized_factors (x : α) : (0 : α) ∉ normalized_factors x := @@ -685,7 +685,7 @@ lemma exists_associated_prime_pow_of_unique_normalized_factor {p r : α} begin use (normalized_factors r).card, have := unique_factorization_monoid.normalized_factors_prod hr, - rwa [multiset.eq_repeat_of_mem (λ b, h), multiset.prod_repeat] at this + rwa [multiset.eq_replicate_of_mem (λ b, h), multiset.prod_replicate] at this end lemma normalized_factors_prod_of_prime [nontrivial α] [unique αˣ] {m : multiset α} @@ -865,9 +865,9 @@ variables [dec_dvd : decidable_rel (has_dvd.dvd : R → R → Prop)] open multiplicity multiset include dec_dvd -lemma le_multiplicity_iff_repeat_le_normalized_factors {a b : R} {n : ℕ} +lemma le_multiplicity_iff_replicate_le_normalized_factors {a b : R} {n : ℕ} (ha : irreducible a) (hb : b ≠ 0) : - ↑n ≤ multiplicity a b ↔ repeat (normalize a) n ≤ normalized_factors b := + ↑n ≤ multiplicity a b ↔ replicate n (normalize a) ≤ normalized_factors b := begin rw ← pow_dvd_iff_le_multiplicity, revert b, @@ -876,12 +876,12 @@ begin split, { rintro ⟨c, rfl⟩, rw [ne.def, pow_succ, mul_assoc, mul_eq_zero, decidable.not_or_iff_and_not] at hb, - rw [pow_succ, mul_assoc, normalized_factors_mul hb.1 hb.2, repeat_succ, + rw [pow_succ, mul_assoc, normalized_factors_mul hb.1 hb.2, replicate_succ, normalized_factors_irreducible ha, singleton_add, cons_le_cons_iff, ← ih hb.2], apply dvd.intro _ rfl }, { rw [multiset.le_iff_exists_add], rintro ⟨u, hu⟩, - rw [← (normalized_factors_prod hb).dvd_iff_dvd_right, hu, prod_add, prod_repeat], + rw [← (normalized_factors_prod hb).dvd_iff_dvd_right, hu, prod_add, prod_replicate], exact (associated.pow_pow $ associated_normalize a).dvd.trans (dvd.intro u.prod rfl) } end @@ -897,9 +897,9 @@ begin apply le_antisymm, { apply part_enat.le_of_lt_add_one, rw [← nat.cast_one, ← nat.cast_add, lt_iff_not_ge, ge_iff_le, - le_multiplicity_iff_repeat_le_normalized_factors ha hb, ← le_count_iff_repeat_le], + le_multiplicity_iff_replicate_le_normalized_factors ha hb, ← le_count_iff_replicate_le], simp }, - rw [le_multiplicity_iff_repeat_le_normalized_factors ha hb, ← le_count_iff_repeat_le], + rw [le_multiplicity_iff_replicate_le_normalized_factors ha hb, ← le_count_iff_replicate_le], end omit dec_dvd @@ -1579,9 +1579,9 @@ eq_of_prod_eq_prod (by rw [factors_prod, factor_set.prod, map_singleton, prod_si subtype.coe_mk]) theorem factors_prime_pow [nontrivial α] {p : associates α} (hp : irreducible p) - (k : ℕ) : factors (p ^ k) = some (multiset.repeat ⟨p, hp⟩ k) := -eq_of_prod_eq_prod (by rw [associates.factors_prod, factor_set.prod, multiset.map_repeat, - multiset.prod_repeat, subtype.coe_mk]) + (k : ℕ) : factors (p ^ k) = some (multiset.replicate k ⟨p, hp⟩) := +eq_of_prod_eq_prod (by rw [associates.factors_prod, factor_set.prod, multiset.map_replicate, + multiset.prod_replicate, subtype.coe_mk]) include dec_irr @@ -1589,7 +1589,7 @@ theorem prime_pow_dvd_iff_le [nontrivial α] {m p : associates α} (h₁ : m ≠ (h₂ : irreducible p) {k : ℕ} : p ^ k ≤ m ↔ k ≤ count p m.factors := begin obtain ⟨a, nz, rfl⟩ := associates.exists_non_zero_rep h₁, - rw [factors_mk _ nz, ← with_top.some_eq_coe, count_some, multiset.le_count_iff_repeat_le, + rw [factors_mk _ nz, ← with_top.some_eq_coe, count_some, multiset.le_count_iff_replicate_le, ← factors_le, factors_prime_pow h₂, factors_mk _ nz], exact with_top.coe_le_coe end diff --git a/src/set_theory/ordinal/fixed_point.lean b/src/set_theory/ordinal/fixed_point.lean index a8c1e5039b5b9..adfdd0d7ec60e 100644 --- a/src/set_theory/ordinal/fixed_point.lean +++ b/src/set_theory/ordinal/fixed_point.lean @@ -344,7 +344,7 @@ 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], + rw [←list.length_replicate n unit.star, ←list.foldr_const f a], apply le_sup }, { rw list.foldr_const f a l, exact le_sup _ _ }, diff --git a/src/tactic/explode.lean b/src/tactic/explode.lean index 4708ff310e253..4307de7d5ced6 100644 --- a/src/tactic/explode.lean +++ b/src/tactic/explode.lean @@ -65,7 +65,7 @@ meta def entries.head (es : entries) : option entry := es.l.head' meta def format_aux : list string → list string → list string → list entry → tactic format | (line :: lines) (dep :: deps) (thm :: thms) (en :: es) := do fmt ← do - { let margin := string.join (list.repeat " │" en.depth), + { let margin := string.join (list.replicate en.depth " │"), let margin := match en.status with | status.sintro := " ├" ++ margin | status.intro := " │" ++ margin ++ " ┌" diff --git a/src/tactic/induction.lean b/src/tactic/induction.lean index 0c3e6d5b12961..c2b6e7cbda8d3 100644 --- a/src/tactic/induction.lean +++ b/src/tactic/induction.lean @@ -1173,7 +1173,7 @@ focus1 $ do let rec_app : name → pexpr := λ rec_suffix, (unchecked_cast expr.mk_app : pexpr → list pexpr → pexpr) (pexpr.mk_explicit (const (iname ++ rec_suffix) [])) - (list.repeat pexpr.mk_placeholder (major_premise_args.length + 1) ++ + (list.replicate (major_premise_args.length + 1) pexpr.mk_placeholder ++ [to_pexpr major_premise]), let rec_suffix := if generate_ihs then "rec_on" else "cases_on", let drec_suffix := if generate_ihs then "drec_on" else "dcases_on", diff --git a/src/tactic/omega/prove_unsats.lean b/src/tactic/omega/prove_unsats.lean index 94233a94f6cf7..ca25905d42176 100644 --- a/src/tactic/omega/prove_unsats.lean +++ b/src/tactic/omega/prove_unsats.lean @@ -22,13 +22,13 @@ meta def prove_neg : int → tactic expr | (int.of_nat _) := failed | -[1+ m] := return `(int.neg_succ_lt_zero %%`(m)) -lemma forall_mem_repeat_zero_eq_zero (m : nat) : - (∀ x ∈ (list.repeat (0 : int) m), x = (0 : int)) := -λ x, list.eq_of_mem_repeat +lemma forall_mem_replicate_zero_eq_zero (m : nat) : + (∀ x ∈ (list.replicate m (0 : int)), x = (0 : int)) := +λ x, list.eq_of_mem_replicate -/-- Return expr of proof that elements of (repeat 0 is.length) are all 0 -/ +/-- Return expr of proof that elements of (replicate is.length 0) are all 0 -/ meta def prove_forall_mem_eq_zero (is : list int) : tactic expr := -return `(forall_mem_repeat_zero_eq_zero is.length) +return `(forall_mem_replicate_zero_eq_zero is.length) /-- Return expr of proof that the combination of linear constraints represented by ks and ts is unsatisfiable -/ diff --git a/src/tactic/solve_by_elim.lean b/src/tactic/solve_by_elim.lean index 71104257a31a1..b51b2f5c7490c 100644 --- a/src/tactic/solve_by_elim.lean +++ b/src/tactic/solve_by_elim.lean @@ -132,7 +132,7 @@ A helper function for trace messages, prepending '....' depending on the current -/ meta def solve_by_elim_trace (n : ℕ) (f : format) : tactic unit := trace_if_enabled `solve_by_elim - (format!"[solve_by_elim {(list.repeat '.' (n+1)).as_string} " ++ f ++ "]") + (format!"[solve_by_elim {(list.replicate (n+1) '.').as_string} " ++ f ++ "]") /-- A helper function to generate trace messages on successful applications. -/ meta def on_success (g : format) (n : ℕ) (e : expr) : tactic unit := diff --git a/src/tactic/unify_equations.lean b/src/tactic/unify_equations.lean index ba29592f4743c..9ac4ad8829a0f 100644 --- a/src/tactic/unify_equations.lean +++ b/src/tactic/unify_equations.lean @@ -140,7 +140,7 @@ do -- Now we generate the actual proof of the target. tgt ← target, - proof ← mk_mapp inj_name (list.repeat none (inj_arity - 3) ++ [some h, some tgt]), + proof ← mk_mapp inj_name (list.replicate (inj_arity - 3) none ++ [some h, some tgt]), eapply proof, (next, ns) ← intron_with num_equations ns base offset, From f15389d94abcd4f2ff4e5e0539135b804698df6c Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Sun, 15 Jan 2023 20:09:54 +0000 Subject: [PATCH 0252/1291] feat(linear_algebra/basis): add basis.ext_elem_iff (#18155) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Replace the lemma ```lean theorem basis.ext_elem (b : basis ι R M) {x y : M} (h : ∀ i, b.repr x i = b.repr y i) : x = y := ``` by an `iff` version: ```lean theorem basis.ext_elem_iff (b : basis ι R M) {x y : M} : x = y ↔ ∀ i, b.repr x i = b.repr y i := ``` --- src/linear_algebra/basis.lean | 12 +++++++----- 1 file changed, 7 insertions(+), 5 deletions(-) diff --git a/src/linear_algebra/basis.lean b/src/linear_algebra/basis.lean index 1fb2f9e25a40e..4baa8870d7ddd 100644 --- a/src/linear_algebra/basis.lean +++ b/src/linear_algebra/basis.lean @@ -234,10 +234,12 @@ by { ext x, omit σ' -/-- Two elements are equal if their coordinates are equal. -/ -theorem ext_elem {x y : M} - (h : ∀ i, b.repr x i = b.repr y i) : x = y := -by { rw [← b.total_repr x, ← b.total_repr y], congr' 1, ext i, exact h i } +/-- Two elements are equal iff their coordinates are equal. -/ +lemma ext_elem_iff {x y : M} : + x = y ↔ (∀ i, b.repr x i = b.repr y i) := +by simp only [← finsupp.ext_iff, embedding_like.apply_eq_iff_eq] + +alias ext_elem_iff ↔ _ _root_.basis.ext_elem lemma repr_eq_iff {b : basis ι R M} {f : M →ₗ[R] ι →₀ R} : ↑b.repr = f ↔ ∀ i, f (b i) = finsupp.single i 1 := @@ -470,7 +472,7 @@ eq_top_iff.mpr $ λ x _, b.mem_span x lemma index_nonempty (b : basis ι R M) [nontrivial M] : nonempty ι := begin obtain ⟨x, y, ne⟩ : ∃ (x y : M), x ≠ y := nontrivial.exists_pair_ne, - obtain ⟨i, _⟩ := not_forall.mp (mt b.ext_elem ne), + obtain ⟨i, _⟩ := not_forall.mp (mt b.ext_elem_iff.2 ne), exact ⟨i⟩ end From b13c1a07b23d09d72cd393a8e72fdd88dab90fa5 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Mon, 16 Jan 2023 01:14:04 +0000 Subject: [PATCH 0253/1291] feat(number_theory/divisors): add `nat.cons_self_proper_divisors` (#18176) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit * Rename `nat.divisors_eq_proper_divisors_insert_self_of_pos` to `nat.insert_self_proper_divisors`, assume `n ≠ 0` instead of `0 < n` and swap LHS with RHS. * Add `nat.cons_self_proper_divisors`. This is easier to use with `finset.prod_cons`/`finset.sum_cons`. --- src/data/finset/basic.lean | 4 ++ src/group_theory/order_of_element.lean | 4 +- src/group_theory/specific_groups/cyclic.lean | 16 ++++---- src/number_theory/divisors.lean | 22 +++++----- .../polynomial/cyclotomic/basic.lean | 41 ++++++++----------- .../polynomial/cyclotomic/eval.lean | 6 +-- 6 files changed, 43 insertions(+), 50 deletions(-) diff --git a/src/data/finset/basic.lean b/src/data/finset/basic.lean index af99e841e5926..69d110ee1a825 100644 --- a/src/data/finset/basic.lean +++ b/src/data/finset/basic.lean @@ -1251,6 +1251,10 @@ theorem erase_insert_of_ne {a b : α} {s : finset α} (h : a ≠ b) : ext $ λ x, have x ≠ b ∧ x = a ↔ x = a, from and_iff_right_of_imp (λ hx, hx.symm ▸ h), by simp only [mem_erase, mem_insert, and_or_distrib_left, this] +theorem erase_cons_of_ne {a b : α} {s : finset α} (ha : a ∉ s) (hb : a ≠ b) : + erase (cons a s ha) b = cons a (erase s b) (λ h, ha $ erase_subset _ _ h) := +by simp only [cons_eq_insert, erase_insert_of_ne hb] + theorem insert_erase {a : α} {s : finset α} (h : a ∈ s) : insert a (erase s a) = s := ext $ assume x, by simp only [mem_insert, mem_erase, or_and_distrib_left, dec_em, true_and]; apply or_iff_right_of_imp; rintro rfl; exact h diff --git a/src/group_theory/order_of_element.lean b/src/group_theory/order_of_element.lean index c8fd227ddc3ed..f7b78b4c442b5 100644 --- a/src/group_theory/order_of_element.lean +++ b/src/group_theory/order_of_element.lean @@ -571,7 +571,7 @@ variables [monoid G] open_locale big_operators @[to_additive sum_card_add_order_of_eq_card_nsmul_eq_zero] -lemma sum_card_order_of_eq_card_pow_eq_one [fintype G] [decidable_eq G] (hn : 0 < n) : +lemma sum_card_order_of_eq_card_pow_eq_one [fintype G] [decidable_eq G] (hn : n ≠ 0) : ∑ m in (finset.range n.succ).filter (∣ n), (finset.univ.filter (λ x : G, order_of x = m)).card = (finset.univ.filter (λ x : G, x ^ n = 1)).card := calc ∑ m in (finset.range n.succ).filter (∣ n), (finset.univ.filter (λ x : G, order_of x = m)).card @@ -581,7 +581,7 @@ calc ∑ m in (finset.range n.succ).filter (∣ n), (finset.univ.filter (λ x : suffices : order_of x ≤ n ∧ order_of x ∣ n ↔ x ^ n = 1, { simpa [nat.lt_succ_iff], }, exact ⟨λ h, let ⟨m, hm⟩ := h.2 in by rw [hm, pow_mul, pow_order_of_eq_one, one_pow], - λ h, ⟨order_of_le_of_pow_eq_one hn h, order_of_dvd_of_pow_eq_one h⟩⟩ + λ h, ⟨order_of_le_of_pow_eq_one hn.bot_lt h, order_of_dvd_of_pow_eq_one h⟩⟩ end)) end finite_monoid diff --git a/src/group_theory/specific_groups/cyclic.lean b/src/group_theory/specific_groups/cyclic.lean index de1bc886463bc..21e961156bfdb 100644 --- a/src/group_theory/specific_groups/cyclic.lean +++ b/src/group_theory/specific_groups/cyclic.lean @@ -297,11 +297,11 @@ private lemma card_order_of_eq_totient_aux₁ : ∀ {d : ℕ}, d ∣ fintype.card α → 0 < (univ.filter (λ a : α, order_of a = d)).card → (univ.filter (λ a : α, order_of a = d)).card = φ d := begin - intros d hd hd0, + intros d hd hpos, induction d using nat.strong_rec' with d IH, - rcases d.eq_zero_or_pos with rfl | hd_pos, + rcases decidable.eq_or_ne d 0 with rfl | hd0, { cases fintype.card_ne_zero (eq_zero_of_zero_dvd hd) }, - rcases card_pos.1 hd0 with ⟨a, ha'⟩, + rcases card_pos.1 hpos with ⟨a, ha'⟩, have ha : order_of a = d := (mem_filter.1 ha').2, have h1 : ∑ m in d.proper_divisors, (univ.filter (λ a : α, order_of a = m)).card = ∑ m in d.proper_divisors, φ m, @@ -309,12 +309,12 @@ begin simp only [mem_filter, mem_range, mem_proper_divisors] at hm, refine IH m hm.2 (hm.1.trans hd) (finset.card_pos.2 ⟨a ^ (d / m), _⟩), simp only [mem_filter, mem_univ, order_of_pow a, ha, true_and, - nat.gcd_eq_right (div_dvd_of_dvd hm.1), nat.div_div_self hm.1 hd_pos.ne'] }, + nat.gcd_eq_right (div_dvd_of_dvd hm.1), nat.div_div_self hm.1 hd0] }, have h2 : ∑ m in d.divisors, (univ.filter (λ a : α, order_of a = m)).card = ∑ m in d.divisors, φ m, - { rw [←filter_dvd_eq_divisors hd_pos.ne', sum_card_order_of_eq_card_pow_eq_one hd_pos, - filter_dvd_eq_divisors hd_pos.ne', sum_totient, ←ha, card_pow_eq_one_eq_order_of_aux hn a] }, - simpa [divisors_eq_proper_divisors_insert_self_of_pos hd_pos, ←h1] using h2, + { rw [←filter_dvd_eq_divisors hd0, sum_card_order_of_eq_card_pow_eq_one hd0, + filter_dvd_eq_divisors hd0, sum_totient, ←ha, card_pow_eq_one_eq_order_of_aux hn a] }, + simpa [← cons_self_proper_divisors hd0, ←h1] using h2, end lemma card_order_of_eq_totient_aux₂ {d : ℕ} (hd : d ∣ fintype.card α) : @@ -328,7 +328,7 @@ begin apply lt_irrefl c, calc c = ∑ m in c.divisors, (univ.filter (λ a : α, order_of a = m)).card : by - { simp only [←filter_dvd_eq_divisors hc0.ne', sum_card_order_of_eq_card_pow_eq_one hc0], + { simp only [←filter_dvd_eq_divisors hc0.ne', sum_card_order_of_eq_card_pow_eq_one hc0.ne'], apply congr_arg card, simp } ... = ∑ m in c.divisors.erase d, (univ.filter (λ a : α, order_of a = m)).card : by diff --git a/src/number_theory/divisors.lean b/src/number_theory/divisors.lean index e239857c3e6ad..3717371862f42 100644 --- a/src/number_theory/divisors.lean +++ b/src/number_theory/divisors.lean @@ -77,10 +77,13 @@ begin simp only [and_comm, ←filter_dvd_eq_proper_divisors hm, mem_filter, mem_range], end -lemma divisors_eq_proper_divisors_insert_self_of_pos (h : 0 < n): - divisors n = has_insert.insert n (proper_divisors n) := -by rw [divisors, proper_divisors, Ico_succ_right_eq_insert_Ico h, finset.filter_insert, - if_pos (dvd_refl n)] +lemma insert_self_proper_divisors (h : n ≠ 0): insert n (proper_divisors n) = divisors n := +by rw [divisors, proper_divisors, Ico_succ_right_eq_insert_Ico (one_le_iff_ne_zero.2 h), + finset.filter_insert, if_pos (dvd_refl n)] + +lemma cons_self_proper_divisors (h : n ≠ 0) : + cons n (proper_divisors n) proper_divisors.not_self_mem = divisors n := +by rw [cons_eq_insert, insert_self_proper_divisors h] @[simp] lemma mem_divisors {m : ℕ} : n ∈ divisors m ↔ (n ∣ m ∧ m ≠ 0) := @@ -245,10 +248,9 @@ end lemma sum_divisors_eq_sum_proper_divisors_add_self : ∑ i in divisors n, i = ∑ i in proper_divisors n, i + n := begin - cases n, + rcases decidable.eq_or_ne n 0 with rfl|hn, { simp }, - { rw [divisors_eq_proper_divisors_insert_self_of_pos (nat.succ_pos _), - finset.sum_insert (proper_divisors.not_self_mem), add_comm] } + { rw [← cons_self_proper_divisors hn, finset.sum_cons, add_comm] } end /-- `n : ℕ` is perfect if and only the sum of the proper divisors of `n` is `n` and `n` @@ -280,8 +282,7 @@ end lemma prime.proper_divisors {p : ℕ} (pp : p.prime) : proper_divisors p = {1} := -by rw [← erase_insert (proper_divisors.not_self_mem), - ← divisors_eq_proper_divisors_insert_self_of_pos pp.pos, +by rw [← erase_insert proper_divisors.not_self_mem, insert_self_proper_divisors pp.ne_zero, pp.divisors, pair_comm, erase_insert (λ con, pp.ne_one (mem_singleton.1 con))] lemma divisors_prime_pow {p : ℕ} (pp : p.prime) (k : ℕ) : @@ -335,8 +336,7 @@ by simp [h.proper_divisors] @[simp, to_additive] lemma prime.prod_divisors {α : Type*} [comm_monoid α] {p : ℕ} {f : ℕ → α} (h : p.prime) : ∏ x in p.divisors, f x = f p * f 1 := -by rw [divisors_eq_proper_divisors_insert_self_of_pos h.pos, - prod_insert proper_divisors.not_self_mem, h.prod_proper_divisors] +by rw [← cons_self_proper_divisors h.ne_zero, prod_cons, h.prod_proper_divisors] lemma proper_divisors_eq_singleton_one_iff_prime : n.proper_divisors = {1} ↔ n.prime := diff --git a/src/ring_theory/polynomial/cyclotomic/basic.lean b/src/ring_theory/polynomial/cyclotomic/basic.lean index c983afb708475..141cc8845711e 100644 --- a/src/ring_theory/polynomial/cyclotomic/basic.lean +++ b/src/ring_theory/polynomial/cyclotomic/basic.lean @@ -180,9 +180,8 @@ lemma cyclotomic'_eq_X_pow_sub_one_div {K : Type*} [comm_ring K] [is_domain K] { (hpos : 0 < n) (h : is_primitive_root ζ n) : cyclotomic' n K = (X ^ n - 1) /ₘ (∏ i in nat.proper_divisors n, cyclotomic' i K) := begin - rw [←prod_cyclotomic'_eq_X_pow_sub_one hpos h, - nat.divisors_eq_proper_divisors_insert_self_of_pos hpos, - finset.prod_insert nat.proper_divisors.not_self_mem], + rw [←prod_cyclotomic'_eq_X_pow_sub_one hpos h, ← nat.cons_self_proper_divisors hpos.ne', + finset.prod_cons], have prod_monic : (∏ i in nat.proper_divisors n, cyclotomic' i K).monic, { apply monic_prod_of_monic, intros i hi, @@ -224,12 +223,9 @@ begin let Q₁ : ℤ[X] := (X ^ k - 1) /ₘ B₁, have huniq : 0 + B * cyclotomic' k K = X ^ k - 1 ∧ (0 : K[X]).degree < B.degree, { split, - { rw [zero_add, mul_comm, ←(prod_cyclotomic'_eq_X_pow_sub_one hpos h), - nat.divisors_eq_proper_divisors_insert_self_of_pos hpos], - simp only [true_and, finset.prod_insert, not_lt, nat.mem_proper_divisors, dvd_refl] }, - rw [degree_zero, bot_lt_iff_ne_bot], - intro habs, - exact (monic.ne_zero Bmo) (degree_eq_bot.1 habs) }, + { rw [zero_add, mul_comm, ← prod_cyclotomic'_eq_X_pow_sub_one hpos h, + ← nat.cons_self_proper_divisors hpos.ne', finset.prod_cons] }, + { simpa only [degree_zero, bot_lt_iff_ne_bot, ne.def, degree_eq_bot] using Bmo.ne_zero } }, replace huniq := div_mod_by_monic_unique (cyclotomic' k K) (0 : K[X]) Bmo huniq, simp only [lifts, ring_hom.mem_srange], use Q₁, @@ -448,8 +444,7 @@ begin convert X_pow_sub_one_mul_prod_cyclotomic_eq_X_pow_sub_one_of_dvd R h using 1, rw mul_assoc, congr' 1, - rw [nat.divisors_eq_proper_divisors_insert_self_of_pos $ pos_of_gt hdn, - finset.insert_sdiff_of_not_mem, finset.prod_insert], + rw [← nat.insert_self_proper_divisors hdn.ne_bot, insert_sdiff_of_not_mem, prod_insert], { exact finset.not_mem_sdiff_of_not_mem_left nat.proper_divisors.not_self_mem }, { exact λ hk, hdn.not_le $ nat.divisor_le hk } end @@ -503,9 +498,8 @@ lemma cyclotomic_eq_X_pow_sub_one_div {R : Type*} [comm_ring R] {n : ℕ} (hpos: 0 < n) : cyclotomic n R = (X ^ n - 1) /ₘ (∏ i in nat.proper_divisors n, cyclotomic i R) := begin nontriviality R, - rw [←prod_cyclotomic_eq_X_pow_sub_one hpos, - nat.divisors_eq_proper_divisors_insert_self_of_pos hpos, - finset.prod_insert nat.proper_divisors.not_self_mem], + rw [←prod_cyclotomic_eq_X_pow_sub_one hpos, ← nat.cons_self_proper_divisors hpos.ne', + finset.prod_cons], have prod_monic : (∏ i in nat.proper_divisors n, cyclotomic i R).monic, { apply monic_prod_of_monic, intros i hi, @@ -670,9 +664,8 @@ lemma eq_cyclotomic_iff {R : Type*} [comm_ring R] {n : ℕ} (hpos: 0 < n) begin nontriviality R, refine ⟨λ hcycl, _, λ hP, _⟩, - { rw [hcycl, ← finset.prod_insert (@nat.proper_divisors.not_self_mem n), - ← nat.divisors_eq_proper_divisors_insert_self_of_pos hpos], - exact prod_cyclotomic_eq_X_pow_sub_one hpos R }, + { rw [hcycl, ← prod_cyclotomic_eq_X_pow_sub_one hpos R, ← nat.cons_self_proper_divisors hpos.ne', + finset.prod_cons] }, { have prod_monic : (∏ i in nat.proper_divisors n, cyclotomic i R).monic, { apply monic_prod_of_monic, intros i hi, @@ -711,7 +704,7 @@ lemma cyclotomic_prime_pow_mul_X_pow_sub_one (R : Type*) [comm_ring R] (p k : by rw [cyclotomic_prime_pow_eq_geom_sum hn.out, geom_sum_mul, ← pow_mul, pow_succ, mul_comm] /-- The constant term of `cyclotomic n R` is `1` if `2 ≤ n`. -/ -lemma cyclotomic_coeff_zero (R : Type*) [comm_ring R] {n : ℕ} (hn : 2 ≤ n) : +lemma cyclotomic_coeff_zero (R : Type*) [comm_ring R] {n : ℕ} (hn : 1 < n) : (cyclotomic n R).coeff 0 = 1 := begin induction n using nat.strong_induction_on with n hi, @@ -732,10 +725,9 @@ begin simp only [finset.prod_const_one] }, simp only [hrw, mul_one, zero_sub, coeff_one_zero, coeff_X_zero, coeff_sub] }, have heq : (X ^ n - 1).coeff 0 = -(cyclotomic n R).coeff 0, - { rw [←prod_cyclotomic_eq_X_pow_sub_one (lt_of_lt_of_le zero_lt_two hn), - nat.divisors_eq_proper_divisors_insert_self_of_pos (lt_of_lt_of_le zero_lt_two hn), - finset.prod_insert nat.proper_divisors.not_self_mem, mul_coeff_zero, coeff_zero_prod, hprod, - mul_neg, mul_one] }, + { rw [← prod_cyclotomic_eq_X_pow_sub_one (zero_le_one.trans_lt hn), + ← nat.cons_self_proper_divisors hn.ne_bot, finset.prod_cons, mul_coeff_zero, + coeff_zero_prod, hprod, mul_neg, mul_one] }, have hzero : (X ^ n - 1).coeff 0 = (-1 : R), { rw coeff_zero_eq_eval_zero _, simp only [zero_pow (lt_of_lt_of_le zero_lt_two hn), eval_X, eval_one, zero_sub, eval_pow, @@ -780,9 +772,8 @@ begin apply units.coe_eq_one.1, simp only [sub_eq_zero.mp hpow, zmod.coe_unit_of_coprime, units.coe_pow] }, rw [is_root.def] at hroot, - rw [← prod_cyclotomic_eq_X_pow_sub_one hpos (zmod p), - nat.divisors_eq_proper_divisors_insert_self_of_pos hpos, - finset.prod_insert nat.proper_divisors.not_self_mem, eval_mul, hroot, zero_mul] + rw [← prod_cyclotomic_eq_X_pow_sub_one hpos (zmod p), ← nat.cons_self_proper_divisors hpos.ne', + finset.prod_cons, eval_mul, hroot, zero_mul] end end order diff --git a/src/ring_theory/polynomial/cyclotomic/eval.lean b/src/ring_theory/polynomial/cyclotomic/eval.lean index 3a5bec45b320a..07d7e226483ec 100644 --- a/src/ring_theory/polynomial/cyclotomic/eval.lean +++ b/src/ring_theory/polynomial/cyclotomic/eval.lean @@ -71,10 +71,8 @@ begin dsimp at ih, have := prod_cyclotomic_eq_geom_sum hn' R, apply_fun eval x at this, - rw [divisors_eq_proper_divisors_insert_self_of_pos hn', finset.erase_insert_of_ne hn''.ne', - finset.prod_insert, eval_mul, eval_geom_sum] at this, - swap, - { simp only [proper_divisors.not_self_mem, mem_erase, and_false, not_false_iff] }, + rw [← cons_self_proper_divisors hn'.ne', finset.erase_cons_of_ne _ hn''.ne', + finset.prod_cons, eval_mul, eval_geom_sum] at this, rcases lt_trichotomy 0 (∑ i in finset.range n, x ^ i) with h | h | h, { apply pos_of_mul_pos_left, { rwa this }, From 47adfab39a11a072db552f47594bf8ed2cf8a722 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Mon, 16 Jan 2023 07:12:55 +0000 Subject: [PATCH 0254/1291] chore(*): sync `list.replicate` with Mathlib 4 (#18181) Sync arguments order and golfs with leanprover-community/mathlib4#1579 --- src/algebra/big_operators/basic.lean | 2 +- src/algebra/big_operators/multiset/basic.lean | 2 +- src/algebra/polynomial/big_operators.lean | 2 +- src/data/list/basic.lean | 42 ++++++++++--------- src/data/list/big_operators/basic.lean | 13 ++---- src/data/list/count.lean | 11 +++-- src/data/list/perm.lean | 8 ++-- src/data/multiset/basic.lean | 28 ++++++------- src/data/nat/digits.lean | 3 +- src/data/nat/factors.lean | 4 +- src/data/polynomial/ring_division.lean | 6 +-- src/group_theory/perm/cycle/type.lean | 2 +- 12 files changed, 62 insertions(+), 61 deletions(-) diff --git a/src/algebra/big_operators/basic.lean b/src/algebra/big_operators/basic.lean index 2ed573801f11f..9a22cfdda19e7 100644 --- a/src/algebra/big_operators/basic.lean +++ b/src/algebra/big_operators/basic.lean @@ -1126,7 +1126,7 @@ begin end @[simp, to_additive] lemma prod_const (b : β) : (∏ x in s, b) = b ^ s.card := -(congr_arg _ $ s.val.map_const b).trans $ multiset.prod_replicate b s.card +(congr_arg _ $ s.val.map_const b).trans $ multiset.prod_replicate s.card b @[to_additive] lemma pow_eq_prod_const (b : β) : ∀ n, b ^ n = ∏ k in range n, b := by simp diff --git a/src/algebra/big_operators/multiset/basic.lean b/src/algebra/big_operators/multiset/basic.lean index 5fa0f421bf34e..f4e1bb72fa95b 100644 --- a/src/algebra/big_operators/multiset/basic.lean +++ b/src/algebra/big_operators/multiset/basic.lean @@ -86,7 +86,7 @@ lemma prod_nsmul (m : multiset α) : ∀ (n : ℕ), (n • m).prod = m.prod ^ n | (n + 1) := by rw [add_nsmul, one_nsmul, pow_add, pow_one, prod_add, prod_nsmul n] -@[simp, to_additive] lemma prod_replicate (a : α) (n : ℕ) : (replicate n a).prod = a ^ n := +@[simp, to_additive] lemma prod_replicate (n : ℕ) (a : α) : (replicate n a).prod = a ^ n := by simp [replicate, list.prod_replicate] @[to_additive] diff --git a/src/algebra/polynomial/big_operators.lean b/src/algebra/polynomial/big_operators.lean index 35f66c3e060ed..f6fdd04502af9 100644 --- a/src/algebra/polynomial/big_operators.lean +++ b/src/algebra/polynomial/big_operators.lean @@ -192,7 +192,7 @@ begin nontriviality R, apply nat_degree_multiset_prod', suffices : (t.map (λ f, leading_coeff f)).prod = 1, { rw this, simp }, - convert prod_replicate (1 : R) t.card, + convert prod_replicate t.card (1 : R), { simp only [eq_replicate, multiset.card_map, eq_self_iff_true, true_and], rintros i hi, obtain ⟨i, hi, rfl⟩ := multiset.mem_map.mp hi, diff --git a/src/data/list/basic.lean b/src/data/list/basic.lean index 93f0a0ec958a6..61268e5fc965b 100644 --- a/src/data/list/basic.lean +++ b/src/data/list/basic.lean @@ -450,13 +450,13 @@ append_inj_right h rfl theorem append_right_cancel {s₁ s₂ t : list α} (h : s₁ ++ t = s₂ ++ t) : s₁ = s₂ := append_inj_left' h rfl -theorem append_right_injective (s : list α) : function.injective (λ t, s ++ t) := +theorem append_right_injective (s : list α) : injective (λ t, s ++ t) := λ t₁ t₂, append_left_cancel theorem append_right_inj {t₁ t₂ : list α} (s) : s ++ t₁ = s ++ t₂ ↔ t₁ = t₂ := (append_right_injective s).eq_iff -theorem append_left_injective (t : list α) : function.injective (λ s, s ++ t) := +theorem append_left_injective (t : list α) : injective (λ s, s ++ t) := λ s₁ s₂, append_right_cancel theorem append_left_inj {s₁ s₂ : list α} (t) : s₁ ++ t = s₂ ++ t ↔ s₁ = s₂ := @@ -475,7 +475,9 @@ end /-! ### replicate -/ +@[simp] theorem replicate_zero (a : α) : replicate 0 a = [] := rfl @[simp] theorem replicate_succ (a : α) (n) : replicate (n + 1) a = a :: replicate n a := rfl +theorem replicate_one (a : α) : replicate 1 a = [a] := rfl @[simp] theorem length_replicate : ∀ n (a : α), length (replicate n a) = n | 0 a := rfl @@ -498,16 +500,19 @@ theorem eq_replicate {a : α} {n} {l : list α} : l = replicate n a ↔ length l ⟨λ h, h.symm ▸ ⟨length_replicate _ _, λ b, eq_of_mem_replicate⟩, λ ⟨e, al⟩, e ▸ eq_replicate_of_mem al⟩ -theorem replicate_add (a : α) (m n) : replicate (m + n) a = replicate m a ++ replicate n a := +theorem replicate_add (m n) (a : α) : replicate (m + n) a = replicate m a ++ replicate n a := by induction m; simp only [*, zero_add, succ_add, replicate]; refl -theorem replicate_subset_singleton (a : α) (n) : replicate n a ⊆ [a] := +theorem replicate_succ' (n) (a : α) : replicate (n + 1) a = replicate n a ++ [a] := +replicate_add n 1 a + +theorem replicate_subset_singleton (n) (a : α) : replicate n a ⊆ [a] := λ b h, mem_singleton.2 (eq_of_mem_replicate h) lemma subset_singleton_iff {a : α} {L : list α} : L ⊆ [a] ↔ ∃ n, L = replicate n a := by simp only [eq_replicate, subset_def, mem_singleton, exists_eq_left'] -@[simp] theorem map_replicate (f : α → β) (a : α) (n) : map f (replicate n a) = replicate n (f a) := +@[simp] theorem map_replicate (f : α → β) (n a) : map f (replicate n a) = replicate n (f a) := by induction n; [refl, simp only [*, replicate, map]]; split; refl @[simp] theorem tail_replicate (n) (a : α) : tail (replicate n a) = replicate (n - 1) a := @@ -516,8 +521,7 @@ by cases n; refl @[simp] theorem join_replicate_nil (n : ℕ) : join (replicate n []) = @nil α := by induction n; [refl, simp only [*, replicate, join, append_nil]] -lemma replicate_right_injective {n : ℕ} (hn : n ≠ 0) : - function.injective (replicate n : α → list α) := +lemma replicate_right_injective {n : ℕ} (hn : n ≠ 0) : injective (replicate n : α → list α) := λ _ _ h, (eq_replicate.1 h).2 _ $ mem_replicate.2 ⟨hn, rfl⟩ lemma replicate_right_inj {a b : α} {n : ℕ} (hn : n ≠ 0) : @@ -529,8 +533,8 @@ lemma replicate_right_inj {a b : α} {n : ℕ} (hn : n ≠ 0) : | 0 := by simp | (n + 1) := (replicate_right_inj n.succ_ne_zero).trans $ by simp only [n.succ_ne_zero, false_or] -lemma replicate_left_injective (a : α) : function.injective (λ n, replicate n a) := -function.left_inverse.injective (λ n, length_replicate n a) +lemma replicate_left_injective (a : α) : injective (λ n, replicate n a) := +left_inverse.injective (λ n, length_replicate n a) @[simp] lemma replicate_left_inj {a : α} {n m : ℕ} : replicate n a = replicate m a ↔ n = m := @@ -660,8 +664,8 @@ by simp only [reverse_core_eq, map_append, map_reverse] by induction l; [refl, simp only [*, reverse_cons, mem_append, mem_singleton, mem_cons_iff, not_mem_nil, false_or, or_false, or_comm]] -@[simp] theorem reverse_replicate (a : α) (n) : reverse (replicate n a) = replicate n a := -eq_replicate.2 ⟨by simp only [length_reverse, length_replicate], +@[simp] theorem reverse_replicate (n) (a : α) : reverse (replicate n a) = replicate n a := +eq_replicate.2 ⟨by rw [length_reverse, length_replicate], λ b h, eq_of_mem_replicate (mem_reverse.1 h)⟩ /-! ### empty -/ @@ -729,12 +733,11 @@ theorem last_mem : ∀ {l : list α} (h : l ≠ []), last l h ∈ l | [a] h := or.inl rfl | (a::b::l) h := or.inr $ by { rw [last_cons_cons], exact last_mem (cons_ne_nil b l) } -lemma last_replicate_succ (a m : ℕ) : +lemma last_replicate_succ (m : ℕ) (a : α) : (replicate (m + 1) a).last (ne_nil_of_length_eq_succ (length_replicate (m + 1) a)) = a := begin - induction m with k IH, - { simp }, - { simpa only [replicate_succ, last] } + simp only [replicate_succ'], + exact last_append_singleton _ end /-! ### last' -/ @@ -1796,12 +1799,13 @@ lemma map_eq_replicate_iff {l : list α} {f : α → β} {b : β} : l.map f = replicate l.length b ↔ (∀ x ∈ l, f x = b) := by simp [eq_replicate] -@[simp] -theorem map_const (l : list α) (b : β) : map (function.const α b) l = replicate l.length b := +@[simp] theorem map_const (l : list α) (b : β) : map (const α b) l = replicate l.length b := map_eq_replicate_iff.mpr (λ x _, rfl) -theorem eq_of_mem_map_const {b₁ b₂ : β} {l : list α} (h : b₁ ∈ map (function.const α b₂) l) : - b₁ = b₂ := +-- Not a `simp` lemma because `function.const` is reducible in Lean 3 +theorem map_const' (l : list α) (b : β) : map (λ _, b) l = replicate l.length b := map_const l b + +theorem eq_of_mem_map_const {b₁ b₂ : β} {l : list α} (h : b₁ ∈ map (const α b₂) l) : b₁ = b₂ := by rw map_const at h; exact eq_of_mem_replicate h /-! ### map₂ -/ diff --git a/src/data/list/big_operators/basic.lean b/src/data/list/big_operators/basic.lean index b17f6051f0435..b49bb59ddf5fa 100644 --- a/src/data/list/big_operators/basic.lean +++ b/src/data/list/big_operators/basic.lean @@ -51,7 +51,7 @@ lemma prod_eq_foldr : l.prod = foldr (*) 1 l := list.rec_on l rfl $ λ a l ihl, by rw [prod_cons, foldr_cons, ihl] @[simp, priority 500, to_additive] -theorem prod_replicate (a : M) (n : ℕ) : (replicate n a).prod = a ^ n := +theorem prod_replicate (n : ℕ) (a : M) : (replicate n a).prod = a ^ n := begin induction n with n ih, { rw pow_zero, refl }, @@ -94,13 +94,8 @@ l.prod_hom₂ (*) mul_mul_mul_comm (mul_one _) _ _ @[simp] lemma prod_map_neg {α} [comm_monoid α] [has_distrib_neg α] (l : list α) : (l.map has_neg.neg).prod = (-1) ^ l.length * l.prod := -begin - convert @prod_map_mul α α _ l (λ _, -1) id, - { ext, rw neg_one_mul, refl }, - { rw [← prod_replicate, map_eq_replicate_iff.2], - exact λ _ _, rfl }, - { rw l.map_id }, -end +by simpa only [id, neg_mul, one_mul, map_const', prod_replicate, map_id] + using @prod_map_mul α α _ l (λ _, -1) id @[to_additive] lemma prod_map_hom (L : list ι) (f : ι → M) {G : Type*} [monoid_hom_class G M N] (g : G) : @@ -478,7 +473,7 @@ lemma prod_map_erase [decidable_eq ι] [comm_monoid M] (f : ι → M) {a} : mul_left_comm (f a) (f b)], } end -@[simp] lemma sum_const_nat (m n : ℕ) : sum (replicate m n) = m * n := +lemma sum_const_nat (m n : ℕ) : sum (replicate m n) = m * n := by rw [sum_replicate, smul_eq_mul] /-- The product of a list of positive natural numbers is positive, diff --git a/src/data/list/count.lean b/src/data/list/count.lean index 2f87dbeb2ae34..5660878757f14 100644 --- a/src/data/list/count.lean +++ b/src/data/list/count.lean @@ -174,12 +174,15 @@ begin exacts [h ▸ count_replicate_self _ _, count_eq_zero_of_not_mem $ mt eq_of_mem_replicate h] end +theorem filter_eq (l : list α) (a : α) : l.filter (eq a) = replicate (count a l) a := +by simp [eq_replicate, count, countp_eq_length_filter, @eq_comm _ _ a] + +theorem filter_eq' (l : list α) (a : α) : l.filter (λ x, x = a) = replicate (count a l) a := +by simp only [filter_eq, @eq_comm _ _ a] + lemma le_count_iff_replicate_sublist {a : α} {l : list α} {n : ℕ} : n ≤ count a l ↔ replicate n a <+ l := -⟨λ h, ((replicate_sublist_replicate a).2 h).trans $ - have filter (eq a) l = replicate (count a l) a, from eq_replicate.2 - ⟨by simp only [count, countp_eq_length_filter], λ b m, (of_mem_filter m).symm⟩, - by rw ← this; apply filter_sublist, +⟨λ h, ((replicate_sublist_replicate a).2 h).trans $ filter_eq l a ▸ filter_sublist _, λ h, by simpa only [count_replicate_self] using h.count_le a⟩ lemma replicate_count_eq_of_count_eq_length {a : α} {l : list α} (h : count a l = length l) : diff --git a/src/data/list/perm.lean b/src/data/list/perm.lean index 557cce5f96c3a..c01819953f2fd 100644 --- a/src/data/list/perm.lean +++ b/src/data/list/perm.lean @@ -150,21 +150,21 @@ theorem perm_cons_append_cons {l l₁ l₂ : list α} (a : α) (p : l ~ l₁++l a::l ~ l₁++(a::l₂) := (p.cons a).trans perm_middle.symm -@[simp] theorem perm_replicate {a : α} {n : ℕ} {l : list α} : +@[simp] theorem perm_replicate {n : ℕ} {a : α} {l : list α} : l ~ replicate n a ↔ l = replicate n a := ⟨λ p, eq_replicate.2 ⟨p.length_eq.trans $ length_replicate _ _, λ b m, eq_of_mem_replicate $ p.subset m⟩, λ h, h ▸ perm.refl _⟩ -@[simp] theorem replicate_perm {a : α} {n : ℕ} {l : list α} : +@[simp] theorem replicate_perm {n : ℕ} {a : α} {l : list α} : replicate n a ~ l ↔ replicate n a = l := (perm_comm.trans perm_replicate).trans eq_comm @[simp] theorem perm_singleton {a : α} {l : list α} : l ~ [a] ↔ l = [a] := -@perm_replicate α a 1 l +@perm_replicate α 1 a l @[simp] theorem singleton_perm {a : α} {l : list α} : [a] ~ l ↔ [a] = l := -@replicate_perm α a 1 l +@replicate_perm α 1 a l alias perm_singleton ↔ perm.eq_singleton _ alias singleton_perm ↔ perm.singleton_eq _ diff --git a/src/data/multiset/basic.lean b/src/data/multiset/basic.lean index a29f9e74da07e..d1c75add6d610 100644 --- a/src/data/multiset/basic.lean +++ b/src/data/multiset/basic.lean @@ -623,20 +623,21 @@ instance is_well_founded_lt : _root_.well_founded_lt (multiset α) := ⟨well_fo /-- `replicate n a` is the multiset containing only `a` with multiplicity `n`. -/ def replicate (n : ℕ) (a : α) : multiset α := replicate n a -lemma coe_replicate (a : α) (n : ℕ) : (list.replicate n a : multiset α) = replicate n a := rfl +lemma coe_replicate (n : ℕ) (a : α) : (list.replicate n a : multiset α) = replicate n a := rfl @[simp] lemma replicate_zero (a : α) : replicate 0 a = 0 := rfl +@[simp] lemma replicate_succ (a : α) (n) : replicate (n + 1) a = a ::ₘ replicate n a := rfl -@[simp] lemma replicate_succ (a : α) (n) : replicate (n+1) a = a ::ₘ replicate n a := rfl - -lemma replicate_add (a : α) (m n : ℕ) : replicate (m + n) a = replicate m a + replicate n a := +lemma replicate_add (m n : ℕ) (a : α) : replicate (m + n) a = replicate m a + replicate n a := congr_arg _ $ list.replicate_add _ _ _ /-- `multiset.replicate` as an `add_monoid_hom`. -/ @[simps] def replicate_add_monoid_hom (a : α) : ℕ →+ multiset α := -{ to_fun := λ n, replicate n a, map_zero' := replicate_zero a, map_add' := replicate_add a } +{ to_fun := λ n, replicate n a, + map_zero' := replicate_zero a, + map_add' := λ _ _, replicate_add _ _ a } -@[simp] lemma replicate_one (a : α) : replicate 1 a = {a} := rfl +lemma replicate_one (a : α) : replicate 1 a = {a} := rfl @[simp] lemma card_replicate : ∀ n (a : α), card (replicate n a) = n := length_replicate @@ -645,7 +646,7 @@ lemma mem_replicate {a b : α} {n : ℕ} : b ∈ replicate n a ↔ n ≠ 0 ∧ b theorem eq_of_mem_replicate {a b : α} {n} : b ∈ replicate n a → b = a := eq_of_mem_replicate theorem eq_replicate_card {a : α} {s : multiset α} : s = replicate s.card a ↔ ∀ b ∈ s, b = a := -quot.induction_on s $ λ l, coe_eq_coe.trans $ perm_replicate.trans $ eq_replicate_length +quot.induction_on s $ λ l, coe_eq_coe.trans $ perm_replicate.trans eq_replicate_length alias eq_replicate_card ↔ _ eq_replicate_of_mem @@ -665,7 +666,7 @@ lemma replicate_right_injective {n : ℕ} (hn : n ≠ 0) : theorem replicate_left_injective (a : α) : function.injective (λ n, replicate n a) := λ m n h, by rw [← (eq_replicate.1 h).1, card_replicate] -theorem replicate_subset_singleton : ∀ (a : α) n, replicate n a ⊆ {a} := replicate_subset_singleton +theorem replicate_subset_singleton : ∀ n (a : α), replicate n a ⊆ {a} := replicate_subset_singleton theorem replicate_le_coe {a : α} {n} {l : list α} : replicate n a ≤ l ↔ list.replicate n a <+ l := @@ -940,6 +941,9 @@ quot.induction_on s $ λ l, congr_arg coe $ map_id _ map (function.const α b) s = replicate s.card b := quot.induction_on s $ λ l, congr_arg coe $ map_const _ _ +-- Not a `simp` lemma because `function.const` is reducibel in Lean 3 +theorem map_const' (s : multiset α) (b : β) : map (λ _, b) s = replicate s.card b := map_const s b + theorem eq_of_mem_map_const {b₁ b₂ : β} {l : list α} (h : b₁ ∈ map (function.const α b₂) l) : b₁ = b₂ := eq_of_mem_replicate $ by rwa map_const at h @@ -1902,16 +1906,12 @@ calc m.attach.count a ... = m.count (a : α) : congr_arg _ m.attach_map_coe lemma filter_eq' (s : multiset α) (b : α) : s.filter (= b) = replicate (count b s) b := -begin - ext a, - rw [count_replicate, count_filter], - exact if_ctx_congr iff.rfl (λ h, congr_arg _ h) (λ h, rfl), -end +quotient.induction_on s $ λ l, congr_arg coe $ filter_eq' l b lemma filter_eq (s : multiset α) (b : α) : s.filter (eq b) = replicate (count b s) b := by simp_rw [←filter_eq', eq_comm] -@[simp] lemma replicate_inter (x : α) (n : ℕ) (s : multiset α) : +@[simp] lemma replicate_inter (n : ℕ) (x : α) (s : multiset α) : replicate n x ∩ s = replicate (min n (s.count x)) x := begin ext y, diff --git a/src/data/nat/digits.lean b/src/data/nat/digits.lean index 60d07ee5c227f..ba1681b9897fb 100644 --- a/src/data/nat/digits.lean +++ b/src/data/nat/digits.lean @@ -347,8 +347,7 @@ begin { cases hm rfl }, { simp } }, { cases m, { cases hm rfl }, - simp_rw [digits_one, list.last_replicate_succ 1 m], - norm_num }, + simpa only [digits_one, list.last_replicate_succ m 1] using one_ne_zero }, revert hm, apply nat.strong_induction_on m, intros n IH hn, diff --git a/src/data/nat/factors.lean b/src/data/nat/factors.lean index 20277b4343685..09b62ef14243c 100644 --- a/src/data/nat/factors.lean +++ b/src/data/nat/factors.lean @@ -158,7 +158,7 @@ lemma prime.factors_pow {p : ℕ} (hp : p.prime) (n : ℕ) : begin symmetry, rw ← list.replicate_perm, - apply nat.factors_unique (list.prod_replicate p n), + apply nat.factors_unique (list.prod_replicate n p), intros q hq, rwa eq_of_mem_replicate hq, end @@ -168,7 +168,7 @@ lemma eq_prime_pow_of_unique_prime_dvd {n p : ℕ} (hpos : n ≠ 0) n = p ^ n.factors.length := begin set k := n.factors.length, - rw [←prod_factors hpos, ←prod_replicate p k, + rw [← prod_factors hpos, ← prod_replicate k p, eq_replicate_of_mem (λ d hd, h (prime_of_mem_factors hd) (dvd_of_mem_factors hd))], end diff --git a/src/data/polynomial/ring_division.lean b/src/data/polynomial/ring_division.lean index cb4c0ce487998..9a44d079bd618 100644 --- a/src/data/polynomial/ring_division.lean +++ b/src/data/polynomial/ring_division.lean @@ -617,9 +617,9 @@ end (s.map (λ a, X - C a)).prod.nat_degree = s.card := begin rw [nat_degree_multiset_prod_of_monic, multiset.map_map], - { convert multiset.sum_replicate 1 _, - { convert multiset.map_const _ 1, ext, apply nat_degree_X_sub_C }, { simp } }, - { intros f hf, obtain ⟨a, ha, rfl⟩ := multiset.mem_map.1 hf, exact monic_X_sub_C a }, + { simp only [(∘), nat_degree_X_sub_C, multiset.map_const, multiset.sum_replicate, smul_eq_mul, + mul_one] }, + { exact multiset.forall_mem_map_iff.2 (λ a _, monic_X_sub_C a) }, end lemma card_roots_X_pow_sub_C {n : ℕ} (hn : 0 < n) (a : R) : diff --git a/src/group_theory/perm/cycle/type.lean b/src/group_theory/perm/cycle/type.lean index cb01ef0cec419..0a8dcbff71766 100644 --- a/src/group_theory/perm/cycle/type.lean +++ b/src/group_theory/perm/cycle/type.lean @@ -482,7 +482,7 @@ begin λ k v, nat.rec (hf1 v).symm (λ k hk, eq.trans (by exact congr_arg σ hk) (hf2 k 1 v)) k, replace hσ : σ ^ (p ^ 1) = 1 := perm.ext (λ v, by rw [pow_one, hσ, hf3, one_apply]), let v₀ : vectors_prod_eq_one G p := - ⟨vector.replicate p 1, (list.prod_replicate 1 p).trans (one_pow p)⟩, + ⟨vector.replicate p 1, (list.prod_replicate p 1).trans (one_pow p)⟩, have hv₀ : σ v₀ = v₀ := subtype.ext (subtype.ext (list.rotate_replicate (1 : G) p 1)), obtain ⟨v, hv1, hv2⟩ := exists_fixed_point_of_prime' Scard hσ hv₀, refine exists_imp_exists (λ g hg, order_of_eq_prime _ (λ hg', hv2 _)) From f2f413b9d4be3a02840d0663dace76e8fe3da053 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Mon, 16 Jan 2023 10:17:32 +0000 Subject: [PATCH 0255/1291] chore(*): add mathlib4 synchronization comments (#18139) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.big_operators.multiset.basic` * `algebra.big_operators.multiset.lemmas` * `algebra.free_monoid.count` * `algebra.module.pointwise_pi` * `algebra.order.lattice_group` * `algebra.order.pointwise` * `data.fin.basic` * `data.fin.succ_pred` * `data.int.range` * `data.list.alist` * `data.list.destutter` * `data.list.nat_antidiagonal` * `data.list.perm` * `data.list.prime` * `data.list.rotate` * `data.list.sigma` * `data.list.sublists` * `data.multiset.basic` * `data.multiset.bind` * `data.multiset.dedup` * `data.multiset.finset_ops` * `data.multiset.fold` * `data.multiset.lattice` * `data.multiset.nat_antidiagonal` * `data.multiset.nodup` * `data.multiset.pi` * `data.multiset.powerset` * `data.multiset.range` * `data.multiset.sections` * `data.multiset.sum` * `data.nat.succ_pred` * `data.pnat.xgcd` * `data.rat.encodable` * `data.real.cau_seq_completion` * `data.set.pointwise.iterate` * `data.set.pointwise.smul` * `data.set.semiring` * `group_theory.group_action.support` * `logic.encodable.basic` * `logic.encodable.lattice` * `order.circular` --- src/algebra/big_operators/multiset/basic.lean | 3 +++ src/algebra/big_operators/multiset/lemmas.lean | 5 ++++- src/algebra/free_monoid/count.lean | 3 +++ src/algebra/module/pointwise_pi.lean | 3 +++ src/algebra/order/lattice_group.lean | 3 +++ src/algebra/order/pointwise.lean | 3 +++ src/data/fin/basic.lean | 3 +++ src/data/fin/succ_pred.lean | 3 +++ src/data/int/range.lean | 3 +++ src/data/list/alist.lean | 3 +++ src/data/list/destutter.lean | 3 +++ src/data/list/nat_antidiagonal.lean | 3 +++ src/data/list/perm.lean | 3 +++ src/data/list/prime.lean | 3 +++ src/data/list/rotate.lean | 3 +++ src/data/list/sigma.lean | 3 +++ src/data/list/sublists.lean | 3 +++ src/data/multiset/basic.lean | 3 +++ src/data/multiset/bind.lean | 3 +++ src/data/multiset/dedup.lean | 3 +++ src/data/multiset/finset_ops.lean | 3 +++ src/data/multiset/fold.lean | 3 +++ src/data/multiset/lattice.lean | 3 +++ src/data/multiset/nat_antidiagonal.lean | 3 +++ src/data/multiset/nodup.lean | 3 +++ src/data/multiset/pi.lean | 3 +++ src/data/multiset/powerset.lean | 3 +++ src/data/multiset/range.lean | 5 ++++- src/data/multiset/sections.lean | 3 +++ src/data/multiset/sum.lean | 3 +++ src/data/nat/succ_pred.lean | 3 +++ src/data/pnat/xgcd.lean | 3 +++ src/data/rat/encodable.lean | 3 +++ src/data/real/cau_seq_completion.lean | 3 +++ src/data/set/pointwise/iterate.lean | 3 +++ src/data/set/pointwise/smul.lean | 3 +++ src/data/set/semiring.lean | 3 +++ src/group_theory/group_action/support.lean | 3 +++ src/logic/encodable/basic.lean | 3 +++ src/logic/encodable/lattice.lean | 3 +++ src/order/circular.lean | 3 +++ 41 files changed, 125 insertions(+), 2 deletions(-) diff --git a/src/algebra/big_operators/multiset/basic.lean b/src/algebra/big_operators/multiset/basic.lean index f4e1bb72fa95b..b159e8d7fb906 100644 --- a/src/algebra/big_operators/multiset/basic.lean +++ b/src/algebra/big_operators/multiset/basic.lean @@ -9,6 +9,9 @@ import data.multiset.basic /-! # Sums and products over multisets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define products and sums indexed by multisets. This is later used to define products and sums indexed by finite sets. diff --git a/src/algebra/big_operators/multiset/lemmas.lean b/src/algebra/big_operators/multiset/lemmas.lean index 8fe2c0c822412..dadbcc8381e52 100644 --- a/src/algebra/big_operators/multiset/lemmas.lean +++ b/src/algebra/big_operators/multiset/lemmas.lean @@ -6,7 +6,10 @@ Authors: Chris Hughes, Bhavik Mehta, Eric Wieser import data.list.big_operators.lemmas import algebra.big_operators.multiset.basic -/-! # Lemmas about `multiset.sum` and `multiset.prod` requiring extra algebra imports -/ +/-! # Lemmas about `multiset.sum` and `multiset.prod` requiring extra algebra imports + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4.-/ variables {ι α β γ : Type*} diff --git a/src/algebra/free_monoid/count.lean b/src/algebra/free_monoid/count.lean index f526f01c6b861..ac642cb2e04ba 100644 --- a/src/algebra/free_monoid/count.lean +++ b/src/algebra/free_monoid/count.lean @@ -9,6 +9,9 @@ import data.list.count /-! # `list.count` as a bundled homomorphism +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define `free_monoid.countp`, `free_monoid.count`, `free_add_monoid.countp`, and `free_add_monoid.count`. These are `list.countp` and `list.count` bundled as multiplicative and additive homomorphisms from `free_monoid` and `free_add_monoid`. diff --git a/src/algebra/module/pointwise_pi.lean b/src/algebra/module/pointwise_pi.lean index 3572a8a1d31a5..3a6335a55f771 100644 --- a/src/algebra/module/pointwise_pi.lean +++ b/src/algebra/module/pointwise_pi.lean @@ -9,6 +9,9 @@ import group_theory.group_action.pi /-! # Pointwise actions on sets in Pi types +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains lemmas about pointwise actions on sets in Pi types. ## Tags diff --git a/src/algebra/order/lattice_group.lean b/src/algebra/order/lattice_group.lean index 2a18a28a237f6..3993c527055ea 100644 --- a/src/algebra/order/lattice_group.lean +++ b/src/algebra/order/lattice_group.lean @@ -10,6 +10,9 @@ import tactic.nth_rewrite /-! # Lattice ordered groups +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Lattice ordered groups were introduced by [Birkhoff][birkhoff1942]. They form the algebraic underpinnings of vector lattices, Banach lattices, AL-space, AM-space etc. diff --git a/src/algebra/order/pointwise.lean b/src/algebra/order/pointwise.lean index c40bd3a229907..f9e5f6392c716 100644 --- a/src/algebra/order/pointwise.lean +++ b/src/algebra/order/pointwise.lean @@ -9,6 +9,9 @@ import data.set.pointwise.smul /-! # Pointwise operations on ordered algebraic objects +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains lemmas about the effect of pointwise operations on sets with an order structure. ## TODO diff --git a/src/data/fin/basic.lean b/src/data/fin/basic.lean index 2be427042b554..a5ba4089bbf05 100644 --- a/src/data/fin/basic.lean +++ b/src/data/fin/basic.lean @@ -12,6 +12,9 @@ import order.hom.set /-! # The finite type with `n` elements +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + `fin n` is the type whose elements are natural numbers smaller than `n`. This file expands on the development in the core library. diff --git a/src/data/fin/succ_pred.lean b/src/data/fin/succ_pred.lean index fc3be9dc186ba..1a8277e5e971f 100644 --- a/src/data/fin/succ_pred.lean +++ b/src/data/fin/succ_pred.lean @@ -9,6 +9,9 @@ import order.succ_pred.basic /-! # Successors and predecessors of `fin n` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we show that `fin n` is both a `succ_order` and a `pred_order`. Note that they are also archimedean, but this is derived from the general instance for well-orderings as opposed to a specific `fin` instance. diff --git a/src/data/int/range.lean b/src/data/int/range.lean index 2bc2fa5cc67fa..9bde57f55dfcd 100644 --- a/src/data/int/range.lean +++ b/src/data/int/range.lean @@ -9,6 +9,9 @@ import data.int.order.basic /-! # Intervals in ℤ +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines integer ranges. `range m n` is the set of integers greater than `m` and strictly less than `n`. diff --git a/src/data/list/alist.lean b/src/data/list/alist.lean index 3b2b912169bea..8648c1c2b183e 100644 --- a/src/data/list/alist.lean +++ b/src/data/list/alist.lean @@ -8,6 +8,9 @@ import data.list.sigma /-! # Association Lists +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines association lists. An association list is a list where every element consists of a key and a value, and no two entries have the same key. The type of the value is allowed to be dependent on the type of the key. diff --git a/src/data/list/destutter.lean b/src/data/list/destutter.lean index 8a761a809ea43..b6a7de18ec97b 100644 --- a/src/data/list/destutter.lean +++ b/src/data/list/destutter.lean @@ -8,6 +8,9 @@ import data.list.chain /-! # Destuttering of Lists +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves theorems about `list.destutter` (in `data.list.defs`), which greedily removes all non-related items that are adjacent in a list, e.g. `[2, 2, 3, 3, 2].destutter (≠) = [2, 3, 2]`. Note that we make no guarantees of being the longest sublist with this property; e.g., diff --git a/src/data/list/nat_antidiagonal.lean b/src/data/list/nat_antidiagonal.lean index de578eb557c0d..c33f04c354aba 100644 --- a/src/data/list/nat_antidiagonal.lean +++ b/src/data/list/nat_antidiagonal.lean @@ -9,6 +9,9 @@ import data.list.range /-! # Antidiagonals in ℕ × ℕ as lists +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the antidiagonals of ℕ × ℕ as lists: the `n`-th antidiagonal is the list of pairs `(i, j)` such that `i + j = n`. This is useful for polynomial multiplication and more generally for sums going from `0` to `n`. diff --git a/src/data/list/perm.lean b/src/data/list/perm.lean index c01819953f2fd..8c1667485d02c 100644 --- a/src/data/list/perm.lean +++ b/src/data/list/perm.lean @@ -11,6 +11,9 @@ import data.nat.factorial.basic /-! # List Permutations +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file introduces the `list.perm` relation, which is true if two lists are permutations of one another. diff --git a/src/data/list/prime.lean b/src/data/list/prime.lean index 277c6e3743e4e..191bd08aa6ba3 100644 --- a/src/data/list/prime.lean +++ b/src/data/list/prime.lean @@ -11,6 +11,9 @@ import data.list.perm /-! # Products of lists of prime elements. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains some theorems relating `prime` and products of `list`s. -/ diff --git a/src/data/list/rotate.lean b/src/data/list/rotate.lean index 9a8d9f7e1f798..451bd4a8fef11 100644 --- a/src/data/list/rotate.lean +++ b/src/data/list/rotate.lean @@ -9,6 +9,9 @@ import data.list.range /-! # List rotation +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves basic results about `list.rotate`, the list rotation. ## Main declarations diff --git a/src/data/list/sigma.lean b/src/data/list/sigma.lean index 600bbc18f7853..10f58ae9d0eca 100644 --- a/src/data/list/sigma.lean +++ b/src/data/list/sigma.lean @@ -9,6 +9,9 @@ import data.list.perm /-! # Utilities for lists of sigmas +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file includes several ways of interacting with `list (sigma β)`, treated as a key-value store. If `α : Type*` and `β : α → Type*`, then we regard `s : sigma β` as having key `s.1 : α` and value diff --git a/src/data/list/sublists.lean b/src/data/list/sublists.lean index b778a4eaa153c..d70adc5c082ee 100644 --- a/src/data/list/sublists.lean +++ b/src/data/list/sublists.lean @@ -8,6 +8,9 @@ import data.list.perm /-! # sublists +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + `list.sublists` gives a list of all (not necessarily contiguous) sublists of a list. This file contains basic results on this function. diff --git a/src/data/multiset/basic.lean b/src/data/multiset/basic.lean index d1c75add6d610..72d9b9e8bf54c 100644 --- a/src/data/multiset/basic.lean +++ b/src/data/multiset/basic.lean @@ -8,6 +8,9 @@ import data.list.perm /-! # Multisets + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. These are implemented as the quotient of a list by permutations. ## Notation We define the global infix notation `::ₘ` for `multiset.cons`. diff --git a/src/data/multiset/bind.lean b/src/data/multiset/bind.lean index 42598b29f4b28..57946c5c01ecb 100644 --- a/src/data/multiset/bind.lean +++ b/src/data/multiset/bind.lean @@ -8,6 +8,9 @@ import algebra.big_operators.multiset.basic /-! # Bind operation for multisets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines a few basic operations on `multiset`, notably the monadic bind. ## Main declarations diff --git a/src/data/multiset/dedup.lean b/src/data/multiset/dedup.lean index dfc2082b24206..d08a5ab927d5b 100644 --- a/src/data/multiset/dedup.lean +++ b/src/data/multiset/dedup.lean @@ -7,6 +7,9 @@ import data.multiset.nodup /-! # Erasing duplicates in a multiset. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ namespace multiset diff --git a/src/data/multiset/finset_ops.lean b/src/data/multiset/finset_ops.lean index e8e5a03ff8a3d..4cc64ca2a3365 100644 --- a/src/data/multiset/finset_ops.lean +++ b/src/data/multiset/finset_ops.lean @@ -8,6 +8,9 @@ import data.multiset.dedup /-! # Preparations for defining operations on `finset`. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The operations here ignore multiplicities, and preparatory for defining the corresponding operations on `finset`. -/ diff --git a/src/data/multiset/fold.lean b/src/data/multiset/fold.lean index 9adcacde43569..fe0fcbbc3e600 100644 --- a/src/data/multiset/fold.lean +++ b/src/data/multiset/fold.lean @@ -8,6 +8,9 @@ import data.list.min_max /-! # The fold operation for a commutative associative operation over a multiset. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ namespace multiset diff --git a/src/data/multiset/lattice.lean b/src/data/multiset/lattice.lean index 1ed7e108374c8..7b279882080d6 100644 --- a/src/data/multiset/lattice.lean +++ b/src/data/multiset/lattice.lean @@ -8,6 +8,9 @@ import data.multiset.fold /-! # Lattice operations on multisets + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ namespace multiset diff --git a/src/data/multiset/nat_antidiagonal.lean b/src/data/multiset/nat_antidiagonal.lean index 40ca75f930169..9b41c72960a56 100644 --- a/src/data/multiset/nat_antidiagonal.lean +++ b/src/data/multiset/nat_antidiagonal.lean @@ -9,6 +9,9 @@ import data.list.nat_antidiagonal /-! # Antidiagonals in ℕ × ℕ as multisets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the antidiagonals of ℕ × ℕ as multisets: the `n`-th antidiagonal is the multiset of pairs `(i, j)` such that `i + j = n`. This is useful for polynomial multiplication and more generally for sums going from `0` to `n`. diff --git a/src/data/multiset/nodup.lean b/src/data/multiset/nodup.lean index d2eb04b019bdc..40ed28dae0c47 100644 --- a/src/data/multiset/nodup.lean +++ b/src/data/multiset/nodup.lean @@ -9,6 +9,9 @@ import data.multiset.range /-! # The `nodup` predicate for multisets without duplicate elements. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ namespace multiset diff --git a/src/data/multiset/pi.lean b/src/data/multiset/pi.lean index 13e166c86da7c..bf0842f7d7290 100644 --- a/src/data/multiset/pi.lean +++ b/src/data/multiset/pi.lean @@ -7,6 +7,9 @@ import data.multiset.nodup /-! # The cartesian product of multisets + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ namespace multiset diff --git a/src/data/multiset/powerset.lean b/src/data/multiset/powerset.lean index 9cf07ef3c2430..57a11693a72e8 100644 --- a/src/data/multiset/powerset.lean +++ b/src/data/multiset/powerset.lean @@ -8,6 +8,9 @@ import data.multiset.nodup /-! # The powerset of a multiset + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ namespace multiset diff --git a/src/data/multiset/range.lean b/src/data/multiset/range.lean index 756ac31bd6d5c..b9959669106ca 100644 --- a/src/data/multiset/range.lean +++ b/src/data/multiset/range.lean @@ -6,7 +6,10 @@ Authors: Mario Carneiro import data.multiset.basic import data.list.range -/-! # `multiset.range n` gives `{0, 1, ..., n-1}` as a multiset. -/ +/-! # `multiset.range n` gives `{0, 1, ..., n-1}` as a multiset. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4.-/ open list nat diff --git a/src/data/multiset/sections.lean b/src/data/multiset/sections.lean index cad402b544edc..6f873a4ca2f40 100644 --- a/src/data/multiset/sections.lean +++ b/src/data/multiset/sections.lean @@ -7,6 +7,9 @@ import data.multiset.bind /-! # Sections of a multiset + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ namespace multiset diff --git a/src/data/multiset/sum.lean b/src/data/multiset/sum.lean index d7c4224f96294..60d03036e5958 100644 --- a/src/data/multiset/sum.lean +++ b/src/data/multiset/sum.lean @@ -8,6 +8,9 @@ import data.multiset.nodup /-! # Disjoint sum of multisets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the disjoint sum of two multisets as `multiset (α ⊕ β)`. Beware not to confuse with the `multiset.sum` operation which computes the additive sum. diff --git a/src/data/nat/succ_pred.lean b/src/data/nat/succ_pred.lean index f07c8d8b6bbb4..676901cd3b4b4 100644 --- a/src/data/nat/succ_pred.lean +++ b/src/data/nat/succ_pred.lean @@ -9,6 +9,9 @@ import order.succ_pred.basic /-! # Successors and predecessors of naturals +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we show that `ℕ` is both an archimedean `succ_order` and an archimedean `pred_order`. -/ diff --git a/src/data/pnat/xgcd.lean b/src/data/pnat/xgcd.lean index 2027de1c747b5..eadc47c825a89 100644 --- a/src/data/pnat/xgcd.lean +++ b/src/data/pnat/xgcd.lean @@ -9,6 +9,9 @@ import data.pnat.prime /-! # Euclidean algorithm for ℕ +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file sets up a version of the Euclidean algorithm that only works with natural numbers. Given `0 < a, b`, it computes the unique `(w, x, y, z, d)` such that the following identities hold: * `a = (w + x) d` diff --git a/src/data/rat/encodable.lean b/src/data/rat/encodable.lean index f846045858aff..7fa5ed6a34ab8 100644 --- a/src/data/rat/encodable.lean +++ b/src/data/rat/encodable.lean @@ -9,6 +9,9 @@ import data.rat.init /-! # The rationals are `encodable`. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + As a consequence we also get the instance `countable ℚ`. This is kept separate from `data.rat.defs` in order to minimize imports. diff --git a/src/data/real/cau_seq_completion.lean b/src/data/real/cau_seq_completion.lean index 0769af1d4a5a5..6a4db83b80edc 100644 --- a/src/data/real/cau_seq_completion.lean +++ b/src/data/real/cau_seq_completion.lean @@ -8,6 +8,9 @@ import data.real.cau_seq /-! # Cauchy completion +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file generalizes the Cauchy completion of `(ℚ, abs)` to the completion of a ring with absolute value. -/ diff --git a/src/data/set/pointwise/iterate.lean b/src/data/set/pointwise/iterate.lean index 9ea81cdc5db2b..c12696ba46bd4 100644 --- a/src/data/set/pointwise/iterate.lean +++ b/src/data/set/pointwise/iterate.lean @@ -9,6 +9,9 @@ import dynamics.fixed_points.basic /-! # Results about pointwise operations on sets with iteration. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ open_locale pointwise diff --git a/src/data/set/pointwise/smul.lean b/src/data/set/pointwise/smul.lean index c1e5817916832..a4a842ee19225 100644 --- a/src/data/set/pointwise/smul.lean +++ b/src/data/set/pointwise/smul.lean @@ -11,6 +11,9 @@ import tactic.by_contra /-! # Pointwise operations of sets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines pointwise algebraic operations on sets. ## Main declarations diff --git a/src/data/set/semiring.lean b/src/data/set/semiring.lean index 8a9825a30d394..22ef765c68d2c 100644 --- a/src/data/set/semiring.lean +++ b/src/data/set/semiring.lean @@ -8,6 +8,9 @@ import data.set.pointwise.smul /-! # Sets as a semiring under union +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines `set_semiring α`, an alias of `set α`, which we endow with `∪` as addition and pointwise `*` as multiplication. If `α` is a (commutative) monoid, `set_semiring α` is a (commutative) semiring. diff --git a/src/group_theory/group_action/support.lean b/src/group_theory/group_action/support.lean index bf39f8da16b97..5d85b3139c86f 100644 --- a/src/group_theory/group_action/support.lean +++ b/src/group_theory/group_action/support.lean @@ -8,6 +8,9 @@ import data.set.pointwise.smul /-! # Support of an element under an action action +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given an action of a group `G` on a type `α`, we say that a set `s : set α` supports an element `a : α` if, for all `g` that fix `s` pointwise, `g` fixes `a`. diff --git a/src/logic/encodable/basic.lean b/src/logic/encodable/basic.lean index 9d47356f8757d..a5d2623a656c0 100644 --- a/src/logic/encodable/basic.lean +++ b/src/logic/encodable/basic.lean @@ -13,6 +13,9 @@ import data.fin.basic /-! # Encodable types +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines encodable (constructively countable) types as a typeclass. This is used to provide explicit encode/decode functions from and to `ℕ`, with the information that those functions are inverses of each other. diff --git a/src/logic/encodable/lattice.lean b/src/logic/encodable/lattice.lean index bbb57d9f9a085..fe4f3fe986991 100644 --- a/src/logic/encodable/lattice.lean +++ b/src/logic/encodable/lattice.lean @@ -9,6 +9,9 @@ import logic.pairwise /-! # Lattice operations on encodable types +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Lemmas about lattice and set operations on encodable types ## Implementation Notes diff --git a/src/order/circular.lean b/src/order/circular.lean index 5255e3f546214..e768116353649 100644 --- a/src/order/circular.lean +++ b/src/order/circular.lean @@ -8,6 +8,9 @@ import data.set.basic /-! # Circular order hierarchy +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines circular preorders, circular partial orders and circular orders. ## Hierarchy From ec2dfcae3677bcdc0d8e906831b1d251dfcbc0f1 Mon Sep 17 00:00:00 2001 From: Lukas Miaskiwskyi Date: Mon, 16 Jan 2023 10:17:33 +0000 Subject: [PATCH 0256/1291] feat(ring_theory/ring_invo): add ring_invo_class (#18175) By its docstring, `ring_equiv_class` yearns to be extended whenever `ring_equiv` is. The `ring_invo` structure fails to heed its call. This file is currently being ported to mathlib4, and this will help, I believe. --- src/ring_theory/ring_invo.lean | 23 +++++++++++++++++++++++ 1 file changed, 23 insertions(+) diff --git a/src/ring_theory/ring_invo.lean b/src/ring_theory/ring_invo.lean index 01324030be5c7..2d8a7bb246d2d 100644 --- a/src/ring_theory/ring_invo.lean +++ b/src/ring_theory/ring_invo.lean @@ -27,13 +27,34 @@ Ring involution variables (R : Type*) +set_option old_structure_cmd true + /-- A ring involution -/ structure ring_invo [semiring R] extends R ≃+* Rᵐᵒᵖ := (involution' : ∀ x, (to_fun (to_fun x).unop).unop = x) +/-- The equivalence of rings underlying a ring involution. -/ +add_decl_doc ring_invo.to_ring_equiv + +/-- `ring_invo_class F R S` states that `F` is a type of ring involutions. +You should extend this class when you extend `ring_invo`. -/ +class ring_invo_class (F : Type*) (R : out_param Type*) [semiring R] + extends ring_equiv_class F R Rᵐᵒᵖ := +(involution : ∀ (f : F) (x), (f (f x).unop).unop = x) + namespace ring_invo variables {R} [semiring R] +instance (R : Type*) [semiring R] : ring_invo_class (ring_invo R) R := +{ coe := to_fun, + inv := inv_fun, + coe_injective' := λ e f h₁ h₂, by { cases e, cases f, congr' }, + map_add := map_add', + map_mul := map_mul', + left_inv := left_inv, + right_inv := right_inv, + involution := involution' } + /-- Construct a ring involution from a ring homomorphism. -/ def mk' (f : R →+* Rᵐᵒᵖ) (involution : ∀ r, (f (f r).unop).unop = r) : ring_invo R := @@ -43,6 +64,8 @@ def mk' (f : R →+* Rᵐᵒᵖ) (involution : ∀ r, (f (f r).unop).unop = r) : involution' := involution, .. f } +/-- Helper instance for when there's too many metavariables to apply +`fun_like.has_coe_to_fun` directly. -/ instance : has_coe_to_fun (ring_invo R) (λ _, R → Rᵐᵒᵖ) := ⟨λ f, f.to_ring_equiv.to_fun⟩ @[simp] From 565eb991e264d0db702722b4bde52ee5173c9950 Mon Sep 17 00:00:00 2001 From: sgouezel Date: Mon, 16 Jan 2023 13:27:53 +0000 Subject: [PATCH 0257/1291] chore(*): remove after the fact `attribute [irreducible]` at several places (#18168) Part of #18164 --- src/algebra/free_algebra.lean | 33 ++-- src/algebra/ring_quot.lean | 64 +++++-- src/data/polynomial/eval.lean | 26 +-- src/linear_algebra/determinant.lean | 7 +- src/linear_algebra/dimension.lean | 11 +- src/linear_algebra/finsupp.lean | 8 +- src/linear_algebra/tensor_algebra/basic.lean | 33 ++-- src/measure_theory/integral/bochner.lean | 177 +++++++++++++----- src/number_theory/padics/ring_homs.lean | 4 +- src/ring_theory/class_group.lean | 6 +- src/ring_theory/fractional_ideal.lean | 40 +++- .../localization/fraction_ring.lean | 12 +- src/ring_theory/polynomial_algebra.lean | 3 +- src/ring_theory/power_basis.lean | 2 - src/topology/algebra/polynomial.lean | 2 +- .../omega_complete_partial_order.lean | 7 +- 16 files changed, 284 insertions(+), 151 deletions(-) diff --git a/src/algebra/free_algebra.lean b/src/algebra/free_algebra.lean index 0dd0cb1a592c4..636e5dbd452f4 100644 --- a/src/algebra/free_algebra.lean +++ b/src/algebra/free_algebra.lean @@ -179,9 +179,10 @@ variables {X} /-- The canonical function `X → free_algebra R X`. -/ -def ι : X → free_algebra R X := λ m, quot.mk _ m +@[irreducible] def ι : X → free_algebra R X := λ m, quot.mk _ m -@[simp] lemma quot_mk_eq_ι (m : X) : quot.mk (free_algebra.rel R X) m = ι R m := rfl +@[simp] lemma quot_mk_eq_ι (m : X) : quot.mk (free_algebra.rel R X) m = ι R m := +by rw [ι] variables {A : Type*} [semiring A] [algebra R A] @@ -230,16 +231,17 @@ private def lift_aux (f : X → A) : (free_algebra R X →ₐ[R] A) := Given a function `f : X → A` where `A` is an `R`-algebra, `lift R f` is the unique lift of `f` to a morphism of `R`-algebras `free_algebra R X → A`. -/ -def lift : (X → A) ≃ (free_algebra R X →ₐ[R] A) := +@[irreducible] def lift : (X → A) ≃ (free_algebra R X →ₐ[R] A) := { to_fun := lift_aux R, inv_fun := λ F, F ∘ (ι R), - left_inv := λ f, by {ext, refl}, + left_inv := λ f, by {ext, rw [ι], refl}, right_inv := λ F, by { ext x, rcases x, induction x, case pre.of : { change ((F : free_algebra R X → A) ∘ (ι R)) _ = _, + rw [ι], refl }, case pre.of_scalar : { change algebra_map _ _ x = F (algebra_map _ _ x), @@ -251,36 +253,35 @@ def lift : (X → A) ≃ (free_algebra R X →ₐ[R] A) := { change lift_aux R (F ∘ ι R) (quot.mk _ _ * quot.mk _ _) = F (quot.mk _ _ * quot.mk _ _), rw [alg_hom.map_mul, alg_hom.map_mul, ha, hb], }, }, } -@[simp] lemma lift_aux_eq (f : X → A) : lift_aux R f = lift R f := rfl +@[simp] lemma lift_aux_eq (f : X → A) : lift_aux R f = lift R f := +by { rw [lift], refl } @[simp] -lemma lift_symm_apply (F : free_algebra R X →ₐ[R] A) : (lift R).symm F = F ∘ (ι R) := rfl +lemma lift_symm_apply (F : free_algebra R X →ₐ[R] A) : (lift R).symm F = F ∘ (ι R) := +by { rw [lift], refl } variables {R X} @[simp] theorem ι_comp_lift (f : X → A) : - (lift R f : free_algebra R X → A) ∘ (ι R) = f := by {ext, refl} + (lift R f : free_algebra R X → A) ∘ (ι R) = f := +by { ext, rw [ι, lift], refl } @[simp] theorem lift_ι_apply (f : X → A) (x) : - lift R f (ι R x) = f x := rfl + lift R f (ι R x) = f x := +by { rw [ι, lift], refl } @[simp] theorem lift_unique (f : X → A) (g : free_algebra R X →ₐ[R] A) : (g : free_algebra R X → A) ∘ (ι R) = f ↔ g = lift R f := -(lift R).symm_apply_eq +by { rw [← (lift R).symm_apply_eq, lift], refl } /-! -At this stage we set the basic definitions as `@[irreducible]`, so from this point onwards one +Since we have set the basic definitions as `@[irreducible]`, from this point onwards one should only use the universal properties of the free algebra, and consider the actual implementation -as a quotient of an inductive type as completely hidden. +as a quotient of an inductive type as completely hidden. -/ -Of course, one still has the option to locally make these definitions `semireducible` if so desired, -and Lean is still willing in some circumstances to do unification based on the underlying -definition. --/ -attribute [irreducible] ι lift -- Marking `free_algebra` irreducible makes `ring` instances inaccessible on quotients. -- https://leanprover.zulipchat.com/#narrow/stream/113488-general/topic/algebra.2Esemiring_to_ring.20breaks.20semimodule.20typeclass.20lookup/near/212580241 -- For now, we avoid this by not marking it irreducible. diff --git a/src/algebra/ring_quot.lean b/src/algebra/ring_quot.lean index 3c0b59d055a7e..231de20a65f7a 100644 --- a/src/algebra/ring_quot.lean +++ b/src/algebra/ring_quot.lean @@ -182,7 +182,7 @@ instance [algebra S R] (r : R → R → Prop) : algebra S (ring_quot r) := /-- The quotient map from a ring to its quotient, as a homomorphism of rings. -/ -def mk_ring_hom (r : R → R → Prop) : R →+* ring_quot r := +@[irreducible] def mk_ring_hom (r : R → R → Prop) : R →+* ring_quot r := { to_fun := λ x, ⟨quot.mk _ x⟩, map_one' := by simp [← one_quot], map_mul' := by simp [mul_quot], @@ -211,7 +211,7 @@ variables {T : Type u₄} [semiring T] Any ring homomorphism `f : R →+* T` which respects a relation `r : R → R → Prop` factors uniquely through a morphism `ring_quot r →+* T`. -/ -def lift {r : R → R → Prop} : +@[irreducible] def lift {r : R → R → Prop} : {f : R →+* T // ∀ ⦃x y⦄, r x y → f x = f y} ≃ (ring_quot r →+* T) := { to_fun := λ f', let f := (f' : R →+* T) in { to_fun := λ x, quot.lift f @@ -228,13 +228,13 @@ def lift {r : R → R → Prop} : map_one' := by simp [← one_quot, f.map_one], map_mul' := by { rintros ⟨⟨x⟩⟩ ⟨⟨y⟩⟩, simp [mul_quot, f.map_mul x y] }, }, inv_fun := λ F, ⟨F.comp (mk_ring_hom r), λ x y h, by { dsimp, rw mk_ring_hom_rel h, }⟩, - left_inv := λ f, by { ext, simp, refl }, - right_inv := λ F, by { ext, simp, refl } } + left_inv := λ f, by { ext, simp [mk_ring_hom] }, + right_inv := λ F, by { ext, simp [mk_ring_hom] } } @[simp] lemma lift_mk_ring_hom_apply (f : R →+* T) {r : R → R → Prop} (w : ∀ ⦃x y⦄, r x y → f x = f y) (x) : lift ⟨f, w⟩ (mk_ring_hom r x) = f x := -rfl +by { simp_rw [lift, mk_ring_hom], refl } -- note this is essentially `lift.symm_apply_eq.mp h` lemma lift_unique (f : R →+* T) {r : R → R → Prop} (w : ∀ ⦃x y⦄, r x y → f x = f y) @@ -243,7 +243,12 @@ by { ext, simp [h], } lemma eq_lift_comp_mk_ring_hom {r : R → R → Prop} (f : ring_quot r →+* T) : f = lift ⟨f.comp (mk_ring_hom r), λ x y h, by { dsimp, rw mk_ring_hom_rel h, }⟩ := -(lift.apply_symm_apply f).symm +begin + conv_lhs { rw ← lift.apply_symm_apply f }, + rw lift, + refl, +end + section comm_ring /-! @@ -261,7 +266,11 @@ lift λ x y h, ideal.quotient.eq.2 $ submodule.mem_Inf.mpr (λ p w, w ⟨x, y, h, sub_add_cancel x y⟩)⟩ @[simp] lemma ring_quot_to_ideal_quotient_apply (r : B → B → Prop) (x : B) : - ring_quot_to_ideal_quotient r (mk_ring_hom r x) = ideal.quotient.mk _ x := rfl + ring_quot_to_ideal_quotient r (mk_ring_hom r x) = ideal.quotient.mk _ x := +begin + simp_rw [ring_quot_to_ideal_quotient, lift, mk_ring_hom], + refl +end /-- The universal ring homomorphism from `B ⧸ ideal.of_rel r` to `ring_quot r`. -/ def ideal_quotient_to_ring_quot (r : B → B → Prop) : @@ -287,7 +296,20 @@ The ring equivalence between `ring_quot r` and `(ideal.of_rel r).quotient` def ring_quot_equiv_ideal_quotient (r : B → B → Prop) : ring_quot r ≃+* B ⧸ ideal.of_rel r := ring_equiv.of_hom_inv (ring_quot_to_ideal_quotient r) (ideal_quotient_to_ring_quot r) - (by { ext, refl, }) (by { ext, refl, }) + (begin + ext, + simp_rw [ring_quot_to_ideal_quotient, lift, mk_ring_hom], + dsimp, + rw [mk_ring_hom], + refl + end) + (begin + ext, + simp_rw [ring_quot_to_ideal_quotient, lift, mk_ring_hom], + dsimp, + rw [mk_ring_hom], + refl + end) end comm_ring @@ -331,20 +353,20 @@ variables (S) /-- The quotient map from an `S`-algebra to its quotient, as a homomorphism of `S`-algebras. -/ -def mk_alg_hom (s : A → A → Prop) : A →ₐ[S] ring_quot s := -{ commutes' := λ r, rfl, +@[irreducible] def mk_alg_hom (s : A → A → Prop) : A →ₐ[S] ring_quot s := +{ commutes' := λ r, by { simp [mk_ring_hom], refl }, ..mk_ring_hom s } @[simp] lemma mk_alg_hom_coe (s : A → A → Prop) : (mk_alg_hom S s : A →+* ring_quot s) = mk_ring_hom s := -rfl +by { simp_rw [mk_alg_hom, mk_ring_hom], refl } lemma mk_alg_hom_rel {s : A → A → Prop} {x y : A} (w : s x y) : mk_alg_hom S s x = mk_alg_hom S s y := by simp [mk_alg_hom, mk_ring_hom, quot.sound (rel.of w)] lemma mk_alg_hom_surjective (s : A → A → Prop) : function.surjective (mk_alg_hom S s) := -by { dsimp [mk_alg_hom], rintro ⟨⟨a⟩⟩, use a, refl, } +by { dsimp [mk_alg_hom, mk_ring_hom], rintro ⟨⟨a⟩⟩, use a, refl, } variables {B : Type u₄} [semiring B] [algebra S B] @@ -361,8 +383,8 @@ end Any `S`-algebra homomorphism `f : A →ₐ[S] B` which respects a relation `s : A → A → Prop` factors uniquely through a morphism `ring_quot s →ₐ[S] B`. -/ -def lift_alg_hom {s : A → A → Prop} : - { f : A →ₐ[S] B // ∀ ⦃x y⦄, s x y → f x = f y} ≃ (ring_quot s →ₐ[S] B) := +@[irreducible] def lift_alg_hom {s : A → A → Prop} : + {f : A →ₐ[S] B // ∀ ⦃x y⦄, s x y → f x = f y} ≃ (ring_quot s →ₐ[S] B) := { to_fun := λ f', let f := (f' : A →ₐ[S] B) in { to_fun := λ x, quot.lift f begin @@ -379,14 +401,14 @@ def lift_alg_hom {s : A → A → Prop} : map_mul' := by { rintros ⟨⟨x⟩⟩ ⟨⟨y⟩⟩, simp [mul_quot, f.map_mul x y], }, commutes' := by { rintros x, simp [← one_quot, smul_quot, algebra.algebra_map_eq_smul_one] } }, inv_fun := λ F, ⟨F.comp (mk_alg_hom S s), λ _ _ h, by { dsimp, erw mk_alg_hom_rel S h }⟩, - left_inv := λ f, by { ext, simp, refl }, - right_inv := λ F, by { ext, simp, refl } } + left_inv := λ f, by { ext, simp [mk_alg_hom, mk_ring_hom] }, + right_inv := λ F, by { ext, simp [mk_alg_hom, mk_ring_hom] } } @[simp] lemma lift_alg_hom_mk_alg_hom_apply (f : A →ₐ[S] B) {s : A → A → Prop} (w : ∀ ⦃x y⦄, s x y → f x = f y) (x) : (lift_alg_hom S ⟨f, w⟩) ((mk_alg_hom S s) x) = f x := -rfl +by { simp_rw [lift_alg_hom, mk_alg_hom, mk_ring_hom], refl, } -- note this is essentially `(lift_alg_hom S).symm_apply_eq.mp h` lemma lift_alg_hom_unique (f : A →ₐ[S] B) {s : A → A → Prop} (w : ∀ ⦃x y⦄, s x y → f x = f y) @@ -395,10 +417,12 @@ by { ext, simp [h], } lemma eq_lift_alg_hom_comp_mk_alg_hom {s : A → A → Prop} (f : ring_quot s →ₐ[S] B) : f = lift_alg_hom S ⟨f.comp (mk_alg_hom S s), λ x y h, by { dsimp, erw mk_alg_hom_rel S h, }⟩ := -((lift_alg_hom S).apply_symm_apply f).symm +begin + conv_lhs { rw ← ((lift_alg_hom S).apply_symm_apply f) }, + rw lift_alg_hom, + refl, +end end algebra -attribute [irreducible] mk_ring_hom mk_alg_hom lift lift_alg_hom - end ring_quot diff --git a/src/data/polynomial/eval.lean b/src/data/polynomial/eval.lean index fb9fe1b2ed7eb..28701b84de5a0 100644 --- a/src/data/polynomial/eval.lean +++ b/src/data/polynomial/eval.lean @@ -31,10 +31,11 @@ variables (f : R →+* S) (x : S) /-- Evaluate a polynomial `p` given a ring hom `f` from the scalar ring to the target and a value `x` for the variable in the target -/ -def eval₂ (p : R[X]) : S := +@[irreducible] def eval₂ (p : R[X]) : S := p.sum (λ e a, f a * x ^ e) -lemma eval₂_eq_sum {f : R →+* S} {x : S} : p.eval₂ f x = p.sum (λ e a, f a * x ^ e) := rfl +lemma eval₂_eq_sum {f : R →+* S} {x : S} : p.eval₂ f x = p.sum (λ e a, f a * x ^ e) := +by rw eval₂ lemma eval₂_congr {R S : Type*} [semiring R] [semiring S] {f g : R →+* S} {s t : S} {φ ψ : R[X]} : @@ -66,7 +67,7 @@ begin end @[simp] lemma eval₂_add : (p + q).eval₂ f x = p.eval₂ f x + q.eval₂ f x := -by { apply sum_add_index; simp [add_mul] } +by { simp only [eval₂_eq_sum], apply sum_add_index; simp [add_mul] } @[simp] lemma eval₂_one : (1 : R[X]).eval₂ f x = 1 := by rw [← C_1, eval₂_C, f.map_one] @@ -256,7 +257,7 @@ variables {x : R} def eval : R → R[X] → R := eval₂ (ring_hom.id _) lemma eval_eq_sum : p.eval x = p.sum (λ e a, a * x ^ e) := -rfl +by { rw [eval, eval₂_eq_sum], refl } lemma eval_eq_sum_range {p : R[X]} (x : R) : p.eval x = ∑ i in finset.range (p.nat_degree + 1), p.coeff i * x ^ i := @@ -382,8 +383,12 @@ lemma is_root.eq_zero (h : is_root p x) : eval x p = 0 := h lemma coeff_zero_eq_eval_zero (p : R[X]) : coeff p 0 = p.eval 0 := calc coeff p 0 = coeff p 0 * 0 ^ 0 : by simp -... = p.eval 0 : eq.symm $ - finset.sum_eq_single _ (λ b _ hb, by simp [zero_pow (nat.pos_of_ne_zero hb)]) (by simp) +... = p.eval 0 : + begin + symmetry, + rw [eval_eq_sum], + exact finset.sum_eq_single _ (λ b _ hb, by simp [zero_pow (nat.pos_of_ne_zero hb)]) (by simp) + end lemma zero_is_root_of_coeff_zero_eq_zero {p : R[X]} (hp : p.coeff 0 = 0) : is_root p 0 := by rwa coeff_zero_eq_eval_zero at hp @@ -401,7 +406,8 @@ section comp /-- The composition of polynomials as a polynomial. -/ def comp (p q : R[X]) : R[X] := p.eval₂ C q -lemma comp_eq_sum_left : p.comp q = p.sum (λ e a, C a * q ^ e) := rfl +lemma comp_eq_sum_left : p.comp q = p.sum (λ e a, C a * q ^ e) := +by rw [comp, eval₂_eq_sum] @[simp] lemma comp_X : p.comp X = p := begin @@ -735,12 +741,10 @@ end end map /-! -After having set up the basic theory of `eval₂`, `eval`, `comp`, and `map`, -we make `eval₂` irreducible. +we have made `eval₂` irreducible from the start. -Perhaps we can make the others irreducible too? +Perhaps we can make also `eval`, `comp`, and `map` irreducible too? -/ -attribute [irreducible] polynomial.eval₂ section hom_eval₂ diff --git a/src/linear_algebra/determinant.lean b/src/linear_algebra/determinant.lean index b670bd9a2e843..a10bd7d34f0f1 100644 --- a/src/linear_algebra/determinant.lean +++ b/src/linear_algebra/determinant.lean @@ -128,7 +128,7 @@ there is no good way to generalize over universe parameters, so we can't fully s type that it does not depend on the choice of basis. Instead you can use the `det_aux_def'` lemma, or avoid mentioning a basis at all using `linear_map.det`. -/ -def det_aux : trunc (basis ι A M) → (M →ₗ[A] M) →* A := +@[irreducible] def det_aux : trunc (basis ι A M) → (M →ₗ[A] M) →* A := trunc.lift (λ b : basis ι A M, (det_monoid_hom).comp (to_matrix_alg_equiv b : (M →ₗ[A] M) →* matrix ι ι A)) @@ -140,10 +140,7 @@ See also `det_aux_def'` which allows you to vary the basis. -/ lemma det_aux_def (b : basis ι A M) (f : M →ₗ[A] M) : linear_map.det_aux (trunc.mk b) f = matrix.det (linear_map.to_matrix b b f) := -rfl - --- Discourage the elaborator from unfolding `det_aux` and producing a huge term. -attribute [irreducible] linear_map.det_aux +by { rw [det_aux], refl } lemma det_aux_def' {ι' : Type*} [fintype ι'] [decidable_eq ι'] (tb : trunc $ basis ι A M) (b' : basis ι' A M) (f : M →ₗ[A] M) : diff --git a/src/linear_algebra/dimension.lean b/src/linear_algebra/dimension.lean index b3710244f1164..7fcadfd6b57f5 100644 --- a/src/linear_algebra/dimension.lean +++ b/src/linear_algebra/dimension.lean @@ -105,7 +105,7 @@ In particular this agrees with the usual notion of the dimension of a vector spa The definition is marked as protected to avoid conflicts with `_root_.rank`, the rank of a linear map. -/ -protected def module.rank : cardinal := +@[irreducible] protected def module.rank : cardinal := ⨆ ι : {s : set V // linear_independent K (coe : s → V)}, #ι.1 end @@ -136,6 +136,7 @@ theorem dim_le {n : ℕ} (H : ∀ s : finset M, linear_independent R (λ i : s, (i : M)) → s.card ≤ n) : module.rank R M ≤ n := begin + rw module.rank, apply csupr_le', rintro ⟨s, li⟩, exact linear_independent_bounded_of_finset_linear_independent_bounded H _ li, @@ -239,7 +240,7 @@ begin apply le_trans, { exact cardinal.lift_mk_le.mpr ⟨(equiv.of_injective _ hv.injective).to_embedding⟩, }, - { simp only [cardinal.lift_le], + { simp only [cardinal.lift_le, module.rank], apply le_trans, swap, exact le_csupr (cardinal.bdd_above_range.{v v} _) ⟨range v, hv.coe_range⟩, @@ -266,6 +267,7 @@ variables (R M) @[simp] lemma dim_punit : module.rank R punit = 0 := begin apply le_bot_iff.mp, + rw module.rank, apply csupr_le', rintro ⟨s, li⟩, apply le_bot_iff.mpr, @@ -775,6 +777,7 @@ theorem basis.mk_eq_dim'' {ι : Type v} (v : basis ι R M) : #ι = module.rank R M := begin haveI := nontrivial_of_invariant_basis_number R, + rw module.rank, apply le_antisymm, { transitivity, swap, @@ -786,10 +789,6 @@ begin apply linear_independent_le_basis v _ li, }, end --- By this stage we want to have a complete API for `module.rank`, --- so we set it `irreducible` here, to keep ourselves honest. -attribute [irreducible] module.rank - theorem basis.mk_range_eq_dim (v : basis ι R M) : #(range v) = module.rank R M := v.reindex_range.mk_eq_dim'' diff --git a/src/linear_algebra/finsupp.lean b/src/linear_algebra/finsupp.lean index 22068faa614a8..f4acf825b69fe 100644 --- a/src/linear_algebra/finsupp.lean +++ b/src/linear_algebra/finsupp.lean @@ -953,14 +953,16 @@ variables (R) Pick some representation of `x : span R w` as a linear combination in `w`, using the axiom of choice. -/ -def span.repr (w : set M) (x : span R w) : w →₀ R := +@[irreducible] def span.repr (w : set M) (x : span R w) : w →₀ R := ((finsupp.mem_span_iff_total _ _ _).mp x.2).some @[simp] lemma span.finsupp_total_repr {w : set M} (x : span R w) : finsupp.total w M R coe (span.repr R w x) = x := -((finsupp.mem_span_iff_total _ _ _).mp x.2).some_spec +begin + rw span.repr, + exact ((finsupp.mem_span_iff_total _ _ _).mp x.2).some_spec +end -attribute [irreducible] span.repr end diff --git a/src/linear_algebra/tensor_algebra/basic.lean b/src/linear_algebra/tensor_algebra/basic.lean index 7ba9ed22644cd..2fe7d35868349 100644 --- a/src/linear_algebra/tensor_algebra/basic.lean +++ b/src/linear_algebra/tensor_algebra/basic.lean @@ -70,47 +70,56 @@ variables {M} /-- The canonical linear map `M →ₗ[R] tensor_algebra R M`. -/ -def ι : M →ₗ[R] (tensor_algebra R M) := +@[irreducible] def ι : M →ₗ[R] (tensor_algebra R M) := { to_fun := λ m, (ring_quot.mk_alg_hom R _ (free_algebra.ι R m)), map_add' := λ x y, by { rw [←alg_hom.map_add], exact ring_quot.mk_alg_hom_rel R rel.add, }, map_smul' := λ r x, by { rw [←alg_hom.map_smul], exact ring_quot.mk_alg_hom_rel R rel.smul, } } lemma ring_quot_mk_alg_hom_free_algebra_ι_eq_ι (m : M) : - ring_quot.mk_alg_hom R (rel R M) (free_algebra.ι R m) = ι R m := rfl + ring_quot.mk_alg_hom R (rel R M) (free_algebra.ι R m) = ι R m := +by { rw [ι], refl } /-- Given a linear map `f : M → A` where `A` is an `R`-algebra, `lift R f` is the unique lift of `f` to a morphism of `R`-algebras `tensor_algebra R M → A`. -/ -@[simps symm_apply] +@[irreducible, simps symm_apply] def lift {A : Type*} [semiring A] [algebra R A] : (M →ₗ[R] A) ≃ (tensor_algebra R M →ₐ[R] A) := { to_fun := ring_quot.lift_alg_hom R ∘ λ f, - ⟨free_algebra.lift R ⇑f, λ x y (h : rel R M x y), by induction h; simp [algebra.smul_def]⟩, + ⟨free_algebra.lift R ⇑f, λ x y (h : rel R M x y), by induction h; + simp only [algebra.smul_def, free_algebra.lift_ι_apply, linear_map.map_smulₛₗ, + ring_hom.id_apply, map_mul, alg_hom.commutes, map_add]⟩, inv_fun := λ F, F.to_linear_map.comp (ι R), - left_inv := λ f, linear_map.ext $ λ x, - (ring_quot.lift_alg_hom_mk_alg_hom_apply _ _ _ _).trans (free_algebra.lift_ι_apply f x), - right_inv := λ F, ring_quot.ring_quot_ext' _ _ _ $ free_algebra.hom_ext $ funext $ λ x, - (ring_quot.lift_alg_hom_mk_alg_hom_apply _ _ _ _).trans (free_algebra.lift_ι_apply _ _) } + left_inv := λ f, begin + rw [ι], + ext1 x, + exact (ring_quot.lift_alg_hom_mk_alg_hom_apply _ _ _ _).trans (free_algebra.lift_ι_apply f x), + end, + right_inv := λ F, ring_quot.ring_quot_ext' _ _ _ $ free_algebra.hom_ext $ funext $ λ x, begin + rw [ι], + exact (ring_quot.lift_alg_hom_mk_alg_hom_apply _ _ _ _).trans (free_algebra.lift_ι_apply _ _) + end } variables {R} @[simp] theorem ι_comp_lift {A : Type*} [semiring A] [algebra R A] (f : M →ₗ[R] A) : - (lift R f).to_linear_map.comp (ι R) = f := (lift R).symm_apply_apply f + (lift R f).to_linear_map.comp (ι R) = f := +by { convert (lift R).symm_apply_apply f, simp only [lift, equiv.coe_fn_symm_mk] } @[simp] theorem lift_ι_apply {A : Type*} [semiring A] [algebra R A] (f : M →ₗ[R] A) (x) : - lift R f (ι R x) = f x := by { dsimp [lift, ι], refl, } + lift R f (ι R x) = f x := +by { conv_rhs { rw ← ι_comp_lift f}, refl } @[simp] theorem lift_unique {A : Type*} [semiring A] [algebra R A] (f : M →ₗ[R] A) (g : tensor_algebra R M →ₐ[R] A) : g.to_linear_map.comp (ι R) = f ↔ g = lift R f := -(lift R).symm_apply_eq +by { rw ← (lift R).symm_apply_eq, simp only [lift, equiv.coe_fn_symm_mk] } -- Marking `tensor_algebra` irreducible makes `ring` instances inaccessible on quotients. -- https://leanprover.zulipchat.com/#narrow/stream/113488-general/topic/algebra.2Esemiring_to_ring.20breaks.20semimodule.20typeclass.20lookup/near/212580241 -- For now, we avoid this by not marking it irreducible. -attribute [irreducible] ι lift @[simp] theorem lift_comp_ι {A : Type*} [semiring A] [algebra R A] (g : tensor_algebra R M →ₐ[R] A) : diff --git a/src/measure_theory/integral/bochner.lean b/src/measure_theory/integral/bochner.lean index a21333f01c18a..5c57984864bb2 100644 --- a/src/measure_theory/integral/bochner.lean +++ b/src/measure_theory/integral/bochner.lean @@ -608,34 +608,54 @@ variables {𝕜} def integral_clm : (α →₁[μ] E) →L[ℝ] E := integral_clm' ℝ /-- The Bochner integral in L1 space -/ -def integral (f : α →₁[μ] E) : E := integral_clm f +@[irreducible] def integral (f : α →₁[μ] E) : E := integral_clm f -lemma integral_eq (f : α →₁[μ] E) : integral f = integral_clm f := rfl +lemma integral_eq (f : α →₁[μ] E) : integral f = integral_clm f := +by simp only [integral] lemma integral_eq_set_to_L1 (f : α →₁[μ] E) : integral f = set_to_L1 (dominated_fin_meas_additive_weighted_smul μ) f := -rfl +by { simp only [integral], refl } @[norm_cast] lemma simple_func.integral_L1_eq_integral (f : α →₁ₛ[μ] E) : integral (f : α →₁[μ] E) = (simple_func.integral f) := -set_to_L1_eq_set_to_L1s_clm (dominated_fin_meas_additive_weighted_smul μ) f +begin + simp only [integral], + exact set_to_L1_eq_set_to_L1s_clm (dominated_fin_meas_additive_weighted_smul μ) f +end variables (α E) @[simp] lemma integral_zero : integral (0 : α →₁[μ] E) = 0 := -map_zero integral_clm +begin + simp only [integral], + exact map_zero integral_clm +end + variables {α E} lemma integral_add (f g : α →₁[μ] E) : integral (f + g) = integral f + integral g := -map_add integral_clm f g +begin + simp only [integral], + exact map_add integral_clm f g +end lemma integral_neg (f : α →₁[μ] E) : integral (-f) = - integral f := -map_neg integral_clm f +begin + simp only [integral], + exact map_neg integral_clm f +end lemma integral_sub (f g : α →₁[μ] E) : integral (f - g) = integral f - integral g := -map_sub integral_clm f g +begin + simp only [integral], + exact map_sub integral_clm f g +end lemma integral_smul (c : 𝕜) (f : α →₁[μ] E) : integral (c • f) = c • integral f := -show (integral_clm' 𝕜) (c • f) = c • (integral_clm' 𝕜) f, from map_smul (integral_clm' 𝕜) c f +begin + simp only [integral], + show (integral_clm' 𝕜) (c • f) = c • (integral_clm' 𝕜) f, from map_smul (integral_clm' 𝕜) c f +end local notation (name := integral_clm) `Integral` := @integral_clm α E _ _ μ _ _ local notation (name := simple_func.integral_clm') `sIntegral` := @@ -645,14 +665,17 @@ lemma norm_Integral_le_one : ‖Integral‖ ≤ 1 := norm_set_to_L1_le (dominated_fin_meas_additive_weighted_smul μ) zero_le_one lemma norm_integral_le (f : α →₁[μ] E) : ‖integral f‖ ≤ ‖f‖ := -calc ‖integral f‖ = ‖Integral f‖ : rfl +calc ‖integral f‖ = ‖Integral f‖ : by simp only [integral] ... ≤ ‖Integral‖ * ‖f‖ : le_op_norm _ _ ... ≤ 1 * ‖f‖ : mul_le_mul_of_nonneg_right norm_Integral_le_one $ norm_nonneg _ ... = ‖f‖ : one_mul _ @[continuity] lemma continuous_integral : continuous (λ (f : α →₁[μ] E), integral f) := -L1.integral_clm.continuous +begin + simp only [integral], + exact L1.integral_clm.continuous +end section pos_part @@ -663,7 +686,8 @@ begin refine @is_closed_property _ _ _ (coe : (α →₁ₛ[μ] ℝ) → (α →₁[μ] ℝ)) (λ f : α →₁[μ] ℝ, integral f = ‖Lp.pos_part f‖ - ‖Lp.neg_part f‖) (simple_func.dense_range one_ne_top) (is_closed_eq _ _) _ f, - { exact cont _ }, + { simp only [integral], + exact cont _ }, { refine continuous.sub (continuous_norm.comp Lp.continuous_pos_part) (continuous_norm.comp Lp.continuous_neg_part) }, -- Show that the property holds for all simple functions in the `L¹` space. @@ -694,7 +718,7 @@ section open_locale classical /-- The Bochner integral -/ -def integral {m : measurable_space α} (μ : measure α) (f : α → E) : E := +@[irreducible] def integral {m : measurable_space α} (μ : measure α) (f : α → E) : E := if hf : integrable f μ then L1.integral (hf.to_L1 f) else 0 end @@ -715,17 +739,20 @@ 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 _ (id _) hf _ _ _ +by { rw [integral], exact @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 := -rfl +by { simp only [integral, L1.integral], refl } 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 +begin + simp only [integral, L1.integral], + exact (L1.set_to_fun_eq_set_to_L1 (dominated_fin_meas_additive_weighted_smul μ) f).symm +end lemma integral_undef (h : ¬ integrable f μ) : ∫ a, f a ∂μ = 0 := -@dif_neg _ (id _) h _ _ _ +by { rw [integral], exact @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 @@ -733,7 +760,10 @@ integral_undef $ not_and_of_not_left _ h variables (α E) lemma integral_zero : ∫ a : α, (0:E) ∂μ = 0 := -set_to_fun_zero (dominated_fin_meas_additive_weighted_smul μ) +begin + simp only [integral, L1.integral], + exact set_to_fun_zero (dominated_fin_meas_additive_weighted_smul μ) +end @[simp] lemma integral_zero' : integral μ (0 : α → E) = 0 := integral_zero α E @@ -742,7 +772,10 @@ variables {α E} lemma integral_add (hf : integrable f μ) (hg : integrable g μ) : ∫ a, f a + g a ∂μ = ∫ a, f a ∂μ + ∫ a, g a ∂μ := -set_to_fun_add (dominated_fin_meas_additive_weighted_smul μ) hf hg +begin + simp only [integral, L1.integral], + exact set_to_fun_add (dominated_fin_meas_additive_weighted_smul μ) hf hg +end lemma integral_add' (hf : integrable f μ) (hg : integrable g μ) : ∫ a, (f + g) a ∂μ = ∫ a, f a ∂μ + ∫ a, g a ∂μ := @@ -750,17 +783,26 @@ integral_add hf hg lemma integral_finset_sum {ι} (s : finset ι) {f : ι → α → E} (hf : ∀ i ∈ s, integrable (f i) μ) : ∫ a, ∑ i in s, f i a ∂μ = ∑ i in s, ∫ a, f i a ∂μ := -set_to_fun_finset_sum (dominated_fin_meas_additive_weighted_smul _) s hf +begin + simp only [integral, L1.integral], + exact set_to_fun_finset_sum (dominated_fin_meas_additive_weighted_smul _) s hf +end lemma integral_neg (f : α → E) : ∫ a, -f a ∂μ = - ∫ a, f a ∂μ := -set_to_fun_neg (dominated_fin_meas_additive_weighted_smul μ) f +begin + simp only [integral, L1.integral], + exact set_to_fun_neg (dominated_fin_meas_additive_weighted_smul μ) f +end lemma integral_neg' (f : α → E) : ∫ a, (-f) a ∂μ = - ∫ a, f a ∂μ := integral_neg f lemma integral_sub (hf : integrable f μ) (hg : integrable g μ) : ∫ a, f a - g a ∂μ = ∫ a, f a ∂μ - ∫ a, g a ∂μ := -set_to_fun_sub (dominated_fin_meas_additive_weighted_smul μ) hf hg +begin + simp only [integral, L1.integral], + exact set_to_fun_sub (dominated_fin_meas_additive_weighted_smul μ) hf hg +end lemma integral_sub' (hf : integrable f μ) (hg : integrable g μ) : ∫ a, (f - g) a ∂μ = ∫ a, f a ∂μ - ∫ a, g a ∂μ := @@ -768,7 +810,10 @@ integral_sub hf hg lemma integral_smul (c : 𝕜) (f : α → E) : ∫ a, c • (f a) ∂μ = c • ∫ a, f a ∂μ := -set_to_fun_smul (dominated_fin_meas_additive_weighted_smul μ) weighted_smul_smul c f +begin + simp only [integral, L1.integral], + exact set_to_fun_smul (dominated_fin_meas_additive_weighted_smul μ) weighted_smul_smul c f +end lemma integral_mul_left {L : Type*} [is_R_or_C L] (r : L) (f : α → L) : ∫ a, r * (f a) ∂μ = r * ∫ a, f a ∂μ := @@ -783,15 +828,24 @@ lemma integral_div {L : Type*} [is_R_or_C L] (r : L) (f : α → L) : by simpa only [←div_eq_mul_inv] using integral_mul_right r⁻¹ f lemma integral_congr_ae (h : f =ᵐ[μ] g) : ∫ a, f a ∂μ = ∫ a, g a ∂μ := -set_to_fun_congr_ae (dominated_fin_meas_additive_weighted_smul μ) h +begin + simp only [integral, L1.integral], + exact set_to_fun_congr_ae (dominated_fin_meas_additive_weighted_smul μ) h +end @[simp] lemma L1.integral_of_fun_eq_integral {f : α → E} (hf : integrable f μ) : ∫ a, (hf.to_L1 f) a ∂μ = ∫ a, f a ∂μ := -set_to_fun_to_L1 (dominated_fin_meas_additive_weighted_smul μ) hf +begin + simp only [integral, L1.integral], + exact set_to_fun_to_L1 (dominated_fin_meas_additive_weighted_smul μ) hf +end @[continuity] lemma continuous_integral : continuous (λ (f : α →₁[μ] E), ∫ a, f a ∂μ) := -continuous_set_to_fun (dominated_fin_meas_additive_weighted_smul μ) +begin + simp only [integral, L1.integral], + exact continuous_set_to_fun (dominated_fin_meas_additive_weighted_smul μ) +end lemma norm_integral_le_lintegral_norm (f : α → E) : ‖∫ a, f a ∂μ‖ ≤ ennreal.to_real (∫⁻ a, (ennreal.of_real ‖f a‖) ∂μ) := @@ -836,7 +890,10 @@ lemma tendsto_integral_of_L1 {ι} (f : α → E) (hfi : integrable f μ) {F : ι → α → E} {l : filter ι} (hFi : ∀ᶠ i in l, integrable (F i) μ) (hF : tendsto (λ i, ∫⁻ x, ‖F i x - f x‖₊ ∂μ) l (𝓝 0)) : tendsto (λ i, ∫ x, F i x ∂μ) l (𝓝 $ ∫ x, f x ∂μ) := -tendsto_set_to_fun_of_L1 (dominated_fin_meas_additive_weighted_smul μ) f hfi hFi hF +begin + simp only [integral, L1.integral], + exact tendsto_set_to_fun_of_L1 (dominated_fin_meas_additive_weighted_smul μ) f hfi hFi hF +end /-- Lebesgue dominated convergence theorem provides sufficient conditions under which almost everywhere convergence of a sequence of functions implies the convergence of their integrals. @@ -849,8 +906,11 @@ theorem tendsto_integral_of_dominated_convergence {F : ℕ → α → E} {f : α (h_bound : ∀ n, ∀ᵐ a ∂μ, ‖F n a‖ ≤ bound a) (h_lim : ∀ᵐ a ∂μ, tendsto (λ n, F n a) at_top (𝓝 (f a))) : tendsto (λn, ∫ a, F n a ∂μ) at_top (𝓝 $ ∫ a, f a ∂μ) := -tendsto_set_to_fun_of_dominated_convergence (dominated_fin_meas_additive_weighted_smul μ) bound - F_measurable bound_integrable h_bound h_lim +begin + simp only [integral, L1.integral], + exact tendsto_set_to_fun_of_dominated_convergence (dominated_fin_meas_additive_weighted_smul μ) + bound F_measurable bound_integrable h_bound h_lim +end /-- Lebesgue dominated convergence theorem for filters with a countable basis -/ lemma tendsto_integral_filter_of_dominated_convergence {ι} {l : filter ι} @@ -861,8 +921,11 @@ lemma tendsto_integral_filter_of_dominated_convergence {ι} {l : filter ι} (bound_integrable : integrable bound μ) (h_lim : ∀ᵐ a ∂μ, tendsto (λ n, F n a) l (𝓝 (f a))) : tendsto (λn, ∫ a, F n a ∂μ) l (𝓝 $ ∫ a, f a ∂μ) := -tendsto_set_to_fun_filter_of_dominated_convergence (dominated_fin_meas_additive_weighted_smul μ) - bound hF_meas h_bound bound_integrable h_lim +begin + simp only [integral, L1.integral], + exact tendsto_set_to_fun_filter_of_dominated_convergence + (dominated_fin_meas_additive_weighted_smul μ) bound hF_meas h_bound bound_integrable h_lim +end /-- Lebesgue dominated convergence theorem for series. -/ lemma has_sum_integral_of_dominated_convergence {ι} [countable ι] @@ -902,16 +965,22 @@ lemma continuous_within_at_of_dominated {F : X → α → E} {x₀ : X} {bound : (bound_integrable : integrable bound μ) (h_cont : ∀ᵐ a ∂μ, continuous_within_at (λ x, F x a) s x₀) : continuous_within_at (λ x, ∫ a, F x a ∂μ) s x₀ := -continuous_within_at_set_to_fun_of_dominated (dominated_fin_meas_additive_weighted_smul μ) hF_meas - h_bound bound_integrable h_cont +begin + simp only [integral, L1.integral], + exact continuous_within_at_set_to_fun_of_dominated (dominated_fin_meas_additive_weighted_smul μ) + hF_meas h_bound bound_integrable h_cont +end lemma continuous_at_of_dominated {F : X → α → E} {x₀ : X} {bound : α → ℝ} (hF_meas : ∀ᶠ x in 𝓝 x₀, ae_strongly_measurable (F x) μ) (h_bound : ∀ᶠ x in 𝓝 x₀, ∀ᵐ a ∂μ, ‖F x a‖ ≤ bound a) (bound_integrable : integrable bound μ) (h_cont : ∀ᵐ a ∂μ, continuous_at (λ x, F x a) x₀) : continuous_at (λ x, ∫ a, F x a ∂μ) x₀ := -continuous_at_set_to_fun_of_dominated (dominated_fin_meas_additive_weighted_smul μ) hF_meas h_bound - bound_integrable h_cont +begin + simp only [integral, L1.integral], + exact continuous_at_set_to_fun_of_dominated (dominated_fin_meas_additive_weighted_smul μ) hF_meas + h_bound bound_integrable h_cont +end lemma continuous_on_of_dominated {F : X → α → E} {bound : α → ℝ} {s : set X} (hF_meas : ∀ x ∈ s, ae_strongly_measurable (F x) μ) @@ -919,15 +988,21 @@ lemma continuous_on_of_dominated {F : X → α → E} {bound : α → ℝ} {s : (bound_integrable : integrable bound μ) (h_cont : ∀ᵐ a ∂μ, continuous_on (λ x, F x a) s) : continuous_on (λ x, ∫ a, F x a ∂μ) s := -continuous_on_set_to_fun_of_dominated (dominated_fin_meas_additive_weighted_smul μ) hF_meas - h_bound bound_integrable h_cont +begin + simp only [integral, L1.integral], + exact continuous_on_set_to_fun_of_dominated (dominated_fin_meas_additive_weighted_smul μ) hF_meas + h_bound bound_integrable h_cont +end lemma continuous_of_dominated {F : X → α → E} {bound : α → ℝ} (hF_meas : ∀ x, ae_strongly_measurable (F x) μ) (h_bound : ∀ x, ∀ᵐ a ∂μ, ‖F x a‖ ≤ bound a) (bound_integrable : integrable bound μ) (h_cont : ∀ᵐ a ∂μ, continuous (λ x, F x a)) : continuous (λ x, ∫ a, F x a ∂μ) := -continuous_set_to_fun_of_dominated (dominated_fin_meas_additive_weighted_smul μ) hF_meas h_bound - bound_integrable h_cont +begin + simp only [integral, L1.integral], + exact continuous_set_to_fun_of_dominated (dominated_fin_meas_additive_weighted_smul μ) hF_meas + h_bound bound_integrable h_cont +end /-- The Bochner integral of a real-valued function `f : α → ℝ` is the difference between the integral of the positive part of `f` and the integral of the negative part of `f`. -/ @@ -1014,8 +1089,11 @@ begin end lemma integral_nonneg_of_ae {f : α → ℝ} (hf : 0 ≤ᵐ[μ] f) : 0 ≤ ∫ a, f a ∂μ := -set_to_fun_nonneg (dominated_fin_meas_additive_weighted_smul μ) - (λ s _ _, weighted_smul_nonneg s) hf +begin + simp only [integral, L1.integral], + exact set_to_fun_nonneg (dominated_fin_meas_additive_weighted_smul μ) + (λ s _ _, weighted_smul_nonneg s) hf +end lemma lintegral_coe_eq_integral (f : α → ℝ≥0) (hfi : integrable (λ x, (f x : ℝ)) μ) : ∫⁻ a, f a ∂μ = ennreal.of_real ∫ a, f a ∂μ := @@ -1136,8 +1214,11 @@ end normed_add_comm_group lemma integral_mono_ae {f g : α → ℝ} (hf : integrable f μ) (hg : integrable g μ) (h : f ≤ᵐ[μ] g) : ∫ a, f a ∂μ ≤ ∫ a, g a ∂μ := -set_to_fun_mono (dominated_fin_meas_additive_weighted_smul μ) (λ s _ _, weighted_smul_nonneg s) - hf hg h +begin + simp only [integral, L1.integral], + exact set_to_fun_mono (dominated_fin_meas_additive_weighted_smul μ) + (λ s _ _, weighted_smul_nonneg s) hf hg h +end @[mono] lemma integral_mono {f g : α → ℝ} (hf : integrable f μ) (hg : integrable g μ) (h : f ≤ g) : ∫ a, f a ∂μ ≤ ∫ a, g a ∂μ := @@ -1200,6 +1281,7 @@ by { rw [← f.integral_eq_integral hfi, simple_func.integral, ← simple_func.i begin cases (@le_top _ _ _ (μ univ)).lt_or_eq with hμ hμ, { haveI : is_finite_measure μ := ⟨hμ⟩, + simp only [integral, L1.integral], exact set_to_fun_const (dominated_fin_meas_additive_weighted_smul _) _, }, { by_cases hc : c = 0, { simp [hc, integral_zero] }, @@ -1223,7 +1305,7 @@ lemma tendsto_integral_approx_on_of_measurable tendsto (λ n, (simple_func.approx_on f hfm s y₀ h₀ n).integral μ) at_top (𝓝 $ ∫ x, f x ∂μ) := begin have hfi' := simple_func.integrable_approx_on hfm hfi h₀ h₀i, - simp only [simple_func.integral_eq_integral _ (hfi' _)], + simp only [simple_func.integral_eq_integral _ (hfi' _), integral, L1.integral], exact tendsto_set_to_fun_approx_on_of_measurable (dominated_fin_meas_additive_weighted_smul μ) hfi hfm hs h₀ h₀i, end @@ -1264,7 +1346,10 @@ end @[simp] lemma integral_zero_measure {m : measurable_space α} (f : α → E) : ∫ x, f x ∂(0 : measure α) = 0 := -set_to_fun_measure_zero (dominated_fin_meas_additive_weighted_smul _) rfl +begin + simp only [integral, L1.integral], + exact set_to_fun_measure_zero (dominated_fin_meas_additive_weighted_smul _) rfl +end theorem integral_finset_sum_measure {ι} {m : measurable_space α} {f : α → E} {μ : ι → measure α} {s : finset ι} (hf : ∀ i ∈ s, integrable f (μ i)) : @@ -1522,8 +1607,6 @@ mk_simp_attribute integral_simps "Simp set for integral rules." attribute [integral_simps] integral_neg integral_smul L1.integral_add L1.integral_sub L1.integral_smul L1.integral_neg -attribute [irreducible] integral L1.integral - section integral_trim variables {H β γ : Type*} [normed_add_comm_group H] diff --git a/src/number_theory/padics/ring_homs.lean b/src/number_theory/padics/ring_homs.lean index e8153eb6910c8..9961c7552ad74 100644 --- a/src/number_theory/padics/ring_homs.lean +++ b/src/number_theory/padics/ring_homs.lean @@ -290,7 +290,7 @@ end /-- `appr n x` gives a value `v : ℕ` such that `x` and `↑v : ℤ_p` are congruent mod `p^n`. See `appr_spec`. -/ -noncomputable def appr : ℤ_[p] → ℕ → ℕ +@[irreducible] noncomputable def appr : ℤ_[p] → ℕ → ℕ | x 0 := 0 | x (n+1) := let y := x - appr x n in @@ -382,8 +382,6 @@ begin exact (dvd_pow_self (p : ℤ_[p]) hc0.ne').mul_left _, }, end -attribute [irreducible] appr - /-- A ring hom from `ℤ_[p]` to `zmod (p^n)`, with underlying function `padic_int.appr n`. -/ def to_zmod_pow (n : ℕ) : ℤ_[p] →+* zmod (p ^ n) := to_zmod_hom (p^n) (λ x, appr x n) diff --git a/src/ring_theory/class_group.lean b/src/ring_theory/class_group.lean index 5e913eba6f342..446a8b7cdbbb8 100644 --- a/src/ring_theory/class_group.lean +++ b/src/ring_theory/class_group.lean @@ -54,17 +54,15 @@ def to_principal_ideal : Kˣ →* (fractional_ideal R⁰ K)ˣ := (by simp only [units.coe_mk, units.coe_mul, span_singleton_mul_span_singleton]), map_one' := ext (by simp only [span_singleton_one, units.coe_mk, units.coe_one]) } -local attribute [semireducible] to_principal_ideal - variables {R K} @[simp] lemma coe_to_principal_ideal (x : Kˣ) : (to_principal_ideal R K x : fractional_ideal R⁰ K) = span_singleton _ x := -rfl +by { simp only [to_principal_ideal], refl } @[simp] lemma to_principal_ideal_eq_iff {I : (fractional_ideal R⁰ K)ˣ} {x : Kˣ} : to_principal_ideal R K x = I ↔ span_singleton R⁰ (x : K) = I := -units.ext_iff +by { simp only [to_principal_ideal], exact units.ext_iff } lemma mem_principal_ideals_iff {I : (fractional_ideal R⁰ K)ˣ} : I ∈ (to_principal_ideal R K).range ↔ ∃ x : K, span_singleton R⁰ x = I := diff --git a/src/ring_theory/fractional_ideal.lean b/src/ring_theory/fractional_ideal.lean index c17ff5cdfe2bc..458ac481a7a8f 100644 --- a/src/ring_theory/fractional_ideal.lean +++ b/src/ring_theory/fractional_ideal.lean @@ -428,31 +428,47 @@ so by making definitions irreducible, we hope to avoid deep unfolds. def mul (I J : fractional_ideal S P) : fractional_ideal S P := ⟨I * J, I.is_fractional.mul J.is_fractional⟩ -local attribute [semireducible] mul +-- local attribute [semireducible] mul instance : has_mul (fractional_ideal S P) := ⟨λ I J, mul I J⟩ @[simp] lemma mul_eq_mul (I J : fractional_ideal S P) : mul I J = I * J := rfl +lemma mul_def (I J : fractional_ideal S P) : I * J = ⟨I * J, I.is_fractional.mul J.is_fractional⟩ := +by simp only [← mul_eq_mul, mul] + @[simp, norm_cast] -lemma coe_mul (I J : fractional_ideal S P) : (↑(I * J) : submodule R P) = I * J := rfl +lemma coe_mul (I J : fractional_ideal S P) : (↑(I * J) : submodule R P) = I * J := +by { simp only [mul_def], refl } @[simp, norm_cast] lemma coe_ideal_mul (I J : ideal R) : (↑(I * J) : fractional_ideal S P) = I * J := -coe_to_submodule_injective $ coe_submodule_mul _ _ _ +begin + simp only [mul_def], + exact coe_to_submodule_injective (coe_submodule_mul _ _ _) +end lemma mul_left_mono (I : fractional_ideal S P) : monotone ((*) I) := -λ J J' h, mul_le.mpr (λ x hx y hy, mul_mem_mul hx (h hy)) +begin + intros J J' h, + simp only [mul_def], + exact mul_le.mpr (λ x hx y hy, mul_mem_mul hx (h hy)) +end lemma mul_right_mono (I : fractional_ideal S P) : monotone (λ J, J * I) := -λ J J' h, mul_le.mpr (λ x hx y hy, mul_mem_mul (h hx) hy) +begin + intros J J' h, + simp only [mul_def], + exact mul_le.mpr (λ x hx y hy, mul_mem_mul (h hx) hy) +end lemma mul_mem_mul {I J : fractional_ideal S P} {i j : P} (hi : i ∈ I) (hj : j ∈ J) : - i * j ∈ I * J := submodule.mul_mem_mul hi hj + i * j ∈ I * J := +by { simp only [mul_def], exact submodule.mul_mem_mul hi hj } lemma mul_le {I J K : fractional_ideal S P} : I * J ≤ K ↔ (∀ (i ∈ I) (j ∈ J), i * j ∈ K) := -submodule.mul_le +by { simp only [mul_def], exact submodule.mul_le } instance : has_pow (fractional_ideal S P) ℕ := ⟨λ I n, ⟨I^n, I.is_fractional.pow n⟩⟩ @@ -464,7 +480,10 @@ lemma coe_pow (I : fractional_ideal S P) (n : ℕ) : ↑(I ^ n) = (I ^ n : submo {C : P → Prop} {r : P} (hr : r ∈ I * J) (hm : ∀ (i ∈ I) (j ∈ J), C (i * j)) (ha : ∀ x y, C x → C y → C (x + y)) : C r := -submodule.mul_induction_on hr hm ha +begin + simp only [mul_def] at hr, + exact submodule.mul_induction_on hr hm ha +end instance : has_nat_cast (fractional_ideal S P) := ⟨nat.unary_cast⟩ @@ -606,7 +625,10 @@ map_coe_ideal g 0 coe_to_submodule_injective (submodule.map_sup _ _ _) @[simp] lemma map_mul : (I * J).map g = I.map g * J.map g := -coe_to_submodule_injective (submodule.map_mul _ _ _) +begin + simp only [mul_def], + exact coe_to_submodule_injective (submodule.map_mul _ _ _) +end @[simp] lemma map_map_symm (g : P ≃ₐ[R] P') : (I.map (g : P →ₐ[R] P')).map (g.symm : P' →ₐ[R] P) = I := diff --git a/src/ring_theory/localization/fraction_ring.lean b/src/ring_theory/localization/fraction_ring.lean index f9b263c6817a5..f726dc02c6f05 100644 --- a/src/ring_theory/localization/fraction_ring.lean +++ b/src/ring_theory/localization/fraction_ring.lean @@ -112,12 +112,10 @@ mk' K ↑(sec (non_zero_divisors A) z).2 mem_non_zero_divisors_iff_ne_zero.2 $ λ h0, h $ eq_zero_of_fst_eq_zero (sec_spec (non_zero_divisors A) z) h0⟩ -local attribute [semireducible] is_fraction_ring.inv - protected lemma mul_inv_cancel (x : K) (hx : x ≠ 0) : x * is_fraction_ring.inv A x = 1 := -show x * dite _ _ _ = 1, begin - rw [dif_neg hx, ←is_unit.mul_left_inj +begin + rw [is_fraction_ring.inv, dif_neg hx, ←is_unit.mul_left_inj (map_units K ⟨(sec _ x).1, mem_non_zero_divisors_iff_ne_zero.2 $ λ h0, hx $ eq_zero_of_fst_eq_zero (sec_spec (non_zero_divisors A) x) h0⟩), one_mul, mul_assoc], @@ -131,7 +129,11 @@ See note [reducible non-instances]. -/ noncomputable def to_field : field K := { inv := is_fraction_ring.inv A, mul_inv_cancel := is_fraction_ring.mul_inv_cancel A, - inv_zero := dif_pos rfl, + inv_zero := begin + change is_fraction_ring.inv A (0 : K) = 0, + rw [is_fraction_ring.inv], + exact dif_pos rfl + end, .. is_fraction_ring.is_domain A, .. show comm_ring K, by apply_instance } diff --git a/src/ring_theory/polynomial_algebra.lean b/src/ring_theory/polynomial_algebra.lean index f41aa6cf97f72..c77c47dce0826 100644 --- a/src/ring_theory/polynomial_algebra.lean +++ b/src/ring_theory/polynomial_algebra.lean @@ -55,8 +55,7 @@ linear_map.to_span_singleton A _ (aeval (polynomial.X : A[X])).to_linear_map lemma to_fun_bilinear_apply_eq_sum (a : A) (p : R[X]) : to_fun_bilinear R A a p = p.sum (λ n r, monomial n (a * algebra_map R A r)) := begin - dsimp [to_fun_bilinear_apply_apply, aeval_def, eval₂_eq_sum, polynomial.sum], - rw finset.smul_sum, + simp only [to_fun_bilinear_apply_apply, aeval_def, eval₂_eq_sum, polynomial.sum, finset.smul_sum], congr' with i : 1, rw [← algebra.smul_def, ←C_mul', mul_smul_comm, C_mul_X_pow_eq_monomial, ←algebra.commutes, ← algebra.smul_def, smul_monomial], diff --git a/src/ring_theory/power_basis.lean b/src/ring_theory/power_basis.lean index e442b65f7bd2e..610f5a9153783 100644 --- a/src/ring_theory/power_basis.lean +++ b/src/ring_theory/power_basis.lean @@ -319,8 +319,6 @@ noncomputable def alg_hom.fintype (pb : power_basis A S) : by letI := classical.dec_eq B; exact fintype.of_equiv _ pb.lift_equiv'.symm -local attribute [irreducible] power_basis.lift - /-- `pb.equiv_of_root pb' h₁ h₂` is an equivalence of algebras with the same power basis, where "the same" means that `pb` is a root of `pb'`s minimal polynomial and vice versa. diff --git a/src/topology/algebra/polynomial.lean b/src/topology/algebra/polynomial.lean index 949efba5a60ef..4467295b0d38c 100644 --- a/src/topology/algebra/polynomial.lean +++ b/src/topology/algebra/polynomial.lean @@ -46,7 +46,7 @@ variables {R S : Type*} [semiring R] [topological_space R] [topological_semiring protected lemma continuous_eval₂ [semiring S] (p : S[X]) (f : S →+* R) : continuous (λ x, p.eval₂ f x) := begin - dsimp only [eval₂_eq_sum, finsupp.sum], + simp only [eval₂_eq_sum, finsupp.sum], exact continuous_finset_sum _ (λ c hc, continuous_const.mul (continuous_pow _)) end diff --git a/src/topology/omega_complete_partial_order.lean b/src/topology/omega_complete_partial_order.lean index 4049204fba2aa..b85acf2d4ae6d 100644 --- a/src/topology/omega_complete_partial_order.lean +++ b/src/topology/omega_complete_partial_order.lean @@ -33,7 +33,6 @@ lemma is_ωSup_iff_is_lub {α : Type u} [preorder α] {c : chain α} {x : α} : by simp [is_ωSup, is_lub, is_least, upper_bounds, lower_bounds] variables (α : Type u) [omega_complete_partial_order α] -local attribute [irreducible] set /-- The characteristic function of open sets is monotone and preserves the limits of chains. -/ @@ -41,8 +40,7 @@ def is_open (s : set α) : Prop := continuous' (λ x, x ∈ s) theorem is_open_univ : is_open α univ := -⟨λ x y h hx, mem_univ _, - by { convert @complete_lattice.top_continuous α Prop _ _, exact rfl }⟩ +⟨λ x y h hx, mem_univ _, @complete_lattice.top_continuous α Prop _ _⟩ theorem is_open.inter (s t : set α) : is_open α s → is_open α t → is_open α (s ∩ t) := complete_lattice.inf_continuous' @@ -55,8 +53,7 @@ begin simp only [Sup_apply, set_of_bijective.surjective.exists, exists_prop, mem_preimage, set_coe.exists, supr_Prop_eq, mem_set_of_eq, subtype.coe_mk, mem_sUnion] }, { intros p hp, - convert hs (set_of p) (mem_preimage.1 hp), - simp only [mem_set_of_eq] }, + exact hs (set_of p) (mem_preimage.1 hp) }, end end Scott From 39afa0b340a56b158991e109cf9d1e5396e67f60 Mon Sep 17 00:00:00 2001 From: sgouezel Date: Mon, 16 Jan 2023 13:27:54 +0000 Subject: [PATCH 0258/1291] chore(*): remove after the fact attribute [irreducible] at several places (2) (#18180) Part of #18164, sequel to #18168. --- src/analysis/inner_product_space/pi_L2.lean | 6 +-- .../inner_product_space/spectrum.lean | 14 ++--- src/data/complex/exponential.lean | 53 +++++++++---------- src/ring_theory/fractional_ideal.lean | 14 ++--- 4 files changed, 42 insertions(+), 45 deletions(-) diff --git a/src/analysis/inner_product_space/pi_L2.lean b/src/analysis/inner_product_space/pi_L2.lean index ab8caa6dc83db..d4ae133a5b30f 100644 --- a/src/analysis/inner_product_space/pi_L2.lean +++ b/src/analysis/inner_product_space/pi_L2.lean @@ -701,7 +701,7 @@ sum has an orthonormal basis indexed by `fin n` and subordinate to that direct s /-- An `n`-dimensional `inner_product_space` equipped with a decomposition as an internal direct sum has an orthonormal basis indexed by `fin n` and subordinate to that direct sum. This function provides the mapping by which it is subordinate. -/ -def direct_sum.is_internal.subordinate_orthonormal_basis_index +@[irreducible] def direct_sum.is_internal.subordinate_orthonormal_basis_index (a : fin n) (hV' : @orthogonal_family 𝕜 _ _ _ _ (λ i, V i) _ (λ i, (V i).subtypeₗᵢ)) : ι := ((hV.sigma_orthonormal_basis_index_equiv hn hV').symm a).1 @@ -712,12 +712,10 @@ lemma direct_sum.is_internal.subordinate_orthonormal_basis_subordinate (hV.subordinate_orthonormal_basis hn hV' a) ∈ V (hV.subordinate_orthonormal_basis_index hn a hV') := by simpa only [direct_sum.is_internal.subordinate_orthonormal_basis, - orthonormal_basis.coe_reindex] + orthonormal_basis.coe_reindex, direct_sum.is_internal.subordinate_orthonormal_basis_index] using hV.collected_orthonormal_basis_mem hV' (λ i, (std_orthonormal_basis 𝕜 (V i))) ((hV.sigma_orthonormal_basis_index_equiv hn hV').symm a) -attribute [irreducible] direct_sum.is_internal.subordinate_orthonormal_basis_index - end subordinate_orthonormal_basis end finite_dimensional diff --git a/src/analysis/inner_product_space/spectrum.lean b/src/analysis/inner_product_space/spectrum.lean index f6d03066c8430..07ce6120ccac8 100644 --- a/src/analysis/inner_product_space/spectrum.lean +++ b/src/analysis/inner_product_space/spectrum.lean @@ -179,7 +179,7 @@ finite-dimensional inner product space `E`. TODO Postcompose with a permutation so that these eigenvectors are listed in increasing order of eigenvalue. -/ -noncomputable def eigenvector_basis : orthonormal_basis (fin n) 𝕜 E := +@[irreducible] noncomputable def eigenvector_basis : orthonormal_basis (fin n) 𝕜 E := hT.direct_sum_is_internal.subordinate_orthonormal_basis hn hT.orthogonal_family_eigenspaces' @@ -187,7 +187,7 @@ hT.direct_sum_is_internal.subordinate_orthonormal_basis hn for a self-adjoint operator `T` on `E`. TODO Postcompose with a permutation so that these eigenvalues are listed in increasing order. -/ -noncomputable def eigenvalues (i : fin n) : ℝ := +@[irreducible] noncomputable def eigenvalues (i : fin n) : ℝ := @is_R_or_C.re 𝕜 _ $ hT.direct_sum_is_internal.subordinate_orthonormal_basis_index hn i hT.orthogonal_family_eigenspaces' @@ -198,11 +198,13 @@ begin let v : E := hT.eigenvector_basis hn i, let μ : 𝕜 := hT.direct_sum_is_internal.subordinate_orthonormal_basis_index hn i hT.orthogonal_family_eigenspaces', + simp_rw [eigenvalues], change has_eigenvector T (is_R_or_C.re μ) v, have key : has_eigenvector T μ v, { have H₁ : v ∈ eigenspace T μ, - { exact hT.direct_sum_is_internal.subordinate_orthonormal_basis_subordinate - hn i hT.orthogonal_family_eigenspaces' }, + { simp_rw [v, eigenvector_basis], + exact hT.direct_sum_is_internal.subordinate_orthonormal_basis_subordinate + hn i hT.orthogonal_family_eigenspaces' }, have H₂ : v ≠ 0 := by simpa using (hT.eigenvector_basis hn).to_basis.ne_zero i, exact ⟨H₁, H₂⟩ }, have re_μ : ↑(is_R_or_C.re μ) = μ, @@ -212,9 +214,7 @@ begin end lemma has_eigenvalue_eigenvalues (i : fin n) : has_eigenvalue T (hT.eigenvalues hn i) := - module.End.has_eigenvalue_of_has_eigenvector (hT.has_eigenvector_eigenvector_basis hn i) - -attribute [irreducible] eigenvector_basis eigenvalues +module.End.has_eigenvalue_of_has_eigenvector (hT.has_eigenvector_eigenvector_basis hn i) @[simp] lemma apply_eigenvector_basis (i : fin n) : T (hT.eigenvector_basis hn i) = (hT.eigenvalues hn i : 𝕜) • hT.eigenvector_basis hn i := diff --git a/src/data/complex/exponential.lean b/src/data/complex/exponential.lean index 21c7ea846b9e9..f41c5375fa3cc 100644 --- a/src/data/complex/exponential.lean +++ b/src/data/complex/exponential.lean @@ -337,7 +337,7 @@ the complex exponential function -/ ⟨λ n, ∑ m in range n, z ^ m / m!, is_cau_exp z⟩ /-- The complex exponential function, defined via its Taylor series -/ -@[pp_nodot] def exp (z : ℂ) : ℂ := lim (exp' z) +@[irreducible, pp_nodot] def exp (z : ℂ) : ℂ := lim (exp' z) /-- The complex sine function, defined via `exp` -/ @[pp_nodot] def sin (z : ℂ) : ℂ := ((exp (-z * I) - exp (z * I)) * I) / 2 @@ -392,8 +392,9 @@ namespace complex variables (x y : ℂ) @[simp] lemma exp_zero : exp 0 = 1 := -lim_eq_of_equiv_const $ - λ ε ε0, ⟨1, λ j hj, begin +begin + rw exp, + refine lim_eq_of_equiv_const (λ ε ε0, ⟨1, λ j hj, _⟩), convert ε0, cases j, { exact absurd hj (not_le_of_gt zero_lt_one) }, @@ -403,33 +404,29 @@ lim_eq_of_equiv_const $ { rw ← ih dec_trivial, simp only [sum_range_succ, pow_succ], simp } } -end⟩ +end lemma exp_add : exp (x + y) = exp x * exp y := -show lim (⟨_, is_cau_exp (x + y)⟩ : cau_seq ℂ abs) = - lim (show cau_seq ℂ abs, from ⟨_, is_cau_exp x⟩) - * lim (show cau_seq ℂ abs, from ⟨_, is_cau_exp y⟩), -from -have hj : ∀ j : ℕ, ∑ m in range j, (x + y) ^ m / m! = - ∑ i in range j, ∑ k in range (i + 1), x ^ k / k! * (y ^ (i - k) / (i - k)!), - from assume j, - finset.sum_congr rfl (λ m hm, begin - rw [add_pow, div_eq_mul_inv, sum_mul], - refine finset.sum_congr rfl (λ i hi, _), - have h₁ : (m.choose i : ℂ) ≠ 0 := nat.cast_ne_zero.2 - (pos_iff_ne_zero.1 (nat.choose_pos (nat.le_of_lt_succ (mem_range.1 hi)))), - have h₂ := nat.choose_mul_factorial_mul_factorial (nat.le_of_lt_succ $ finset.mem_range.1 hi), - rw [← h₂, nat.cast_mul, nat.cast_mul, mul_inv, mul_inv], - simp only [mul_left_comm (m.choose i : ℂ), mul_assoc, mul_left_comm (m.choose i : ℂ)⁻¹, - mul_comm (m.choose i : ℂ)], - rw inv_mul_cancel h₁, - simp [div_eq_mul_inv, mul_comm, mul_assoc, mul_left_comm] - end), -by rw lim_mul_lim; - exact eq.symm (lim_eq_lim_of_equiv (by dsimp; simp only [hj]; - exact cauchy_product (is_cau_abs_exp x) (is_cau_exp y))) - -attribute [irreducible] complex.exp +begin + have hj : ∀ j : ℕ, ∑ m in range j, (x + y) ^ m / m! = + ∑ i in range j, ∑ k in range (i + 1), x ^ k / k! * (y ^ (i - k) / (i - k)!), + { assume j, + refine finset.sum_congr rfl (λ m hm, _), + rw [add_pow, div_eq_mul_inv, sum_mul], + refine finset.sum_congr rfl (λ i hi, _), + have h₁ : (m.choose i : ℂ) ≠ 0 := nat.cast_ne_zero.2 + (pos_iff_ne_zero.1 (nat.choose_pos (nat.le_of_lt_succ (mem_range.1 hi)))), + have h₂ := nat.choose_mul_factorial_mul_factorial (nat.le_of_lt_succ $ finset.mem_range.1 hi), + rw [← h₂, nat.cast_mul, nat.cast_mul, mul_inv, mul_inv], + simp only [mul_left_comm (m.choose i : ℂ), mul_assoc, mul_left_comm (m.choose i : ℂ)⁻¹, + mul_comm (m.choose i : ℂ)], + rw inv_mul_cancel h₁, + simp [div_eq_mul_inv, mul_comm, mul_assoc, mul_left_comm] }, + simp_rw [exp, exp', lim_mul_lim], + apply (lim_eq_lim_of_equiv _).symm, + simp only [hj], + exact cauchy_product (is_cau_abs_exp x) (is_cau_exp y) +end lemma exp_list_sum (l : list ℂ) : exp l.sum = (l.map exp).prod := @monoid_hom.map_list_prod (multiplicative ℂ) ℂ _ _ ⟨exp, exp_zero, exp_add⟩ l diff --git a/src/ring_theory/fractional_ideal.lean b/src/ring_theory/fractional_ideal.lean index 458ac481a7a8f..7a28d8888942c 100644 --- a/src/ring_theory/fractional_ideal.lean +++ b/src/ring_theory/fractional_ideal.lean @@ -1094,14 +1094,15 @@ variables (S) def span_singleton (x : P) : fractional_ideal S P := ⟨span R {x}, is_fractional_span_singleton x⟩ -local attribute [semireducible] span_singleton +-- local attribute [semireducible] span_singleton @[simp] lemma coe_span_singleton (x : P) : - (span_singleton S x : submodule R P) = span R {x} := rfl + (span_singleton S x : submodule R P) = span R {x} := +by { rw span_singleton, refl } @[simp] lemma mem_span_singleton {x y : P} : x ∈ span_singleton S y ↔ ∃ (z : R), z • y = x := -submodule.mem_span_singleton +by { rw span_singleton, exact submodule.mem_span_singleton } lemma mem_span_singleton_self (x : P) : x ∈ span_singleton S x := @@ -1111,12 +1112,13 @@ variables {S} lemma span_singleton_eq_span_singleton [no_zero_smul_divisors R P] {x y : P} : span_singleton S x = span_singleton S y ↔ ∃ z : Rˣ, z • x = y := -by { rw [← submodule.span_singleton_eq_span_singleton], exact subtype.mk_eq_mk } +by { rw [← submodule.span_singleton_eq_span_singleton, span_singleton, span_singleton], + exact subtype.mk_eq_mk } lemma eq_span_singleton_of_principal (I : fractional_ideal S P) [is_principal (I : submodule R P)] : I = span_singleton S (generator (I : submodule R P)) := -coe_to_submodule_injective (span_singleton_generator ↑I).symm +by { rw span_singleton, exact coe_to_submodule_injective (span_singleton_generator ↑I).symm } lemma is_principal_iff (I : fractional_ideal S P) : is_principal (I : submodule R P) ↔ ∃ x, I = span_singleton S x := @@ -1301,7 +1303,7 @@ begin obtain ⟨a, aI, -, ha⟩ := exists_eq_span_singleton_mul I, use (algebra_map R K a)⁻¹ * algebra_map R K (generator aI), suffices : I = span_singleton R⁰ ((algebra_map R K a)⁻¹ * algebra_map R K (generator aI)), - { exact congr_arg subtype.val this }, + { rw span_singleton at this, exact congr_arg subtype.val this }, conv_lhs { rw [ha, ←span_singleton_generator aI] }, rw [ideal.submodule_span_eq, coe_ideal_span_singleton (generator aI), span_singleton_mul_span_singleton] From f4ecb599422baaf39055d8278c7d9ef3b5b72b88 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Mon, 16 Jan 2023 16:42:04 +0000 Subject: [PATCH 0259/1291] feat(data/sum/basic): add some missing lemmas (#18184) Forward-ported to Mathlib 4 in leanprover-community/mathlib4#1583 --- src/data/sum/basic.lean | 16 +++++++++++++--- 1 file changed, 13 insertions(+), 3 deletions(-) diff --git a/src/data/sum/basic.lean b/src/data/sum/basic.lean index a935d2eec34f2..8436c31ff01de 100644 --- a/src/data/sum/basic.lean +++ b/src/data/sum/basic.lean @@ -82,12 +82,18 @@ section get variables {x y : α ⊕ β} -lemma get_left_eq_none_iff : x.get_left = none ↔ x.is_right := +@[simp] lemma get_left_eq_none_iff : x.get_left = none ↔ x.is_right := by cases x; simp only [get_left, is_right, coe_sort_tt, coe_sort_ff, eq_self_iff_true] -lemma get_right_eq_none_iff : x.get_right = none ↔ x.is_left := +@[simp] lemma get_right_eq_none_iff : x.get_right = none ↔ x.is_left := by cases x; simp only [get_right, is_left, coe_sort_tt, coe_sort_ff, eq_self_iff_true] +@[simp] lemma get_left_eq_some_iff {a} : x.get_left = some a ↔ x = inl a := +by cases x; simp only [get_left] + +@[simp] lemma get_right_eq_some_iff {b} : x.get_right = some b ↔ x = inr b := +by cases x; simp only [get_right] + @[simp] lemma bnot_is_left (x : α ⊕ β) : bnot x.is_left = x.is_right := by cases x; refl @[simp] lemma is_left_eq_ff : x.is_left = ff ↔ x.is_right := by cases x; simp lemma not_is_left : ¬x.is_left ↔ x.is_right := by simp @@ -156,9 +162,13 @@ funext $ map_map f' g' f g @[simp] lemma map_id_id (α β) : sum.map (@id α) (@id β) = id := funext $ λ x, sum.rec_on x (λ _, rfl) (λ _, rfl) +lemma elim_map {α β γ δ ε : Sort*} {f₁ : α → β} {f₂ : β → ε} {g₁ : γ → δ} {g₂ : δ → ε} {x} : + sum.elim f₂ g₂ (sum.map f₁ g₁ x) = sum.elim (f₂ ∘ f₁) (g₂ ∘ g₁) x := +by cases x; refl + lemma elim_comp_map {α β γ δ ε : Sort*} {f₁ : α → β} {f₂ : β → ε} {g₁ : γ → δ} {g₂ : δ → ε} : sum.elim f₂ g₂ ∘ sum.map f₁ g₁ = sum.elim (f₂ ∘ f₁) (g₂ ∘ g₁) := -by ext (_|_); refl +funext $ λ _, elim_map @[simp] lemma is_left_map (f : α → β) (g : γ → δ) (x : α ⊕ γ) : is_left (x.map f g) = is_left x := From 91288e351d51b3f0748f0a38faa7613fb0ae2ada Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 16 Jan 2023 20:04:36 +0000 Subject: [PATCH 0260/1291] feat(data/fin/tuple): rename `fin.append` to `matrix.vec_append`, introduce a new `fin.append` and `fin.repeat`. (#10346) We already had `fin.append v w h`, which combines the append operation with casting. This commit removes the `h` argument, allowing it to be defeq to `fin.add_cases`, and moves the previous definition to the name `matrix.vec_append` matching `matrix.vec_cons` and similar. Another advantage of dropping `h` is that it is clearer how to state lemmas like `append_assoc`, as we only have one proof argument to keep track of instead of four. As it turns out, to implement a `gmonoid` structure on tuples (under concatenation), `fin.append` without the `h` argument is all that's needed. We implement `matrix.vec_append` in terms of `fin.append`, but provide a lemma that unfolds it to the old definition so as to avoid having to rewrite all the other proofs. Removing `matrix.vec_append` entirely is left to investigate in some future PR. Co-authored-by: ChrisHughes24 --- src/data/fin/tuple/basic.lean | 117 ++++++++++++++++-- src/data/fin/vec_notation.lean | 58 ++++++--- src/linear_algebra/cross_product.lean | 12 +- src/model_theory/encoding.lean | 4 +- .../legendre_symbol/gauss_sum.lean | 2 +- src/order/jordan_holder.lean | 76 ++++++------ 6 files changed, 196 insertions(+), 73 deletions(-) diff --git a/src/data/fin/tuple/basic.lean b/src/data/fin/tuple/basic.lean index 290eab4b23d51..b388460072fc0 100644 --- a/src/data/fin/tuple/basic.lean +++ b/src/data/fin/tuple/basic.lean @@ -25,6 +25,8 @@ We define the following operations: * `fin.insert_nth` : insert an element to a tuple at a given position. * `fin.find p` : returns the first index `n` where `p n` is satisfied, and `none` if it is never satisfied. +* `fin.append a b` : append two tuples. +* `fin.repeat n a` : repeat a tuple `n` times. -/ universes u v @@ -246,17 +248,110 @@ set.ext $ λ y, exists_fin_succ.trans $ eq_comm.or iff.rfl set.range (fin.cons x b : fin n.succ → α) = insert x (set.range b) := by rw [range_fin_succ, cons_zero, tail_cons] -/-- `fin.append ho u v` appends two vectors of lengths `m` and `n` to produce -one of length `o = m + n`. `ho` provides control of definitional equality -for the vector length. -/ -def append {α : Type*} {o : ℕ} (ho : o = m + n) (u : fin m → α) (v : fin n → α) : fin o → α := -λ i, if h : (i : ℕ) < m - then u ⟨i, h⟩ - else v ⟨(i : ℕ) - m, (tsub_lt_iff_left (le_of_not_lt h)).2 (ho ▸ i.property)⟩ - -@[simp] lemma fin_append_apply_zero {α : Type*} {o : ℕ} (ho : (o + 1) = (m + 1) + n) - (u : fin (m + 1) → α) (v : fin n → α) : - fin.append ho u v 0 = u 0 := rfl +section append + +/-- Append a tuple of length `m` to a tuple of length `n` to get a tuple of length `m + n`. +This is a non-dependent version of `fin.add_cases`. -/ +def append {α : Type*} (a : fin m → α) (b : fin n → α) : fin (m + n) → α := +@fin.add_cases _ _ (λ _, α) a b + +@[simp] lemma append_left {α : Type*} (u : fin m → α) (v : fin n → α) (i : fin m) : + append u v (fin.cast_add n i) = u i := +add_cases_left _ _ _ + +@[simp] lemma append_right {α : Type*} (u : fin m → α) (v : fin n → α) (i : fin n) : + append u v (nat_add m i) = v i := +add_cases_right _ _ _ + +lemma append_right_nil {α : Type*} (u : fin m → α) (v : fin n → α) (hv : n = 0) : + append u v = u ∘ fin.cast (by rw [hv, add_zero]) := +begin + refine funext (fin.add_cases (λ l, _) (λ r, _)), + { rw [append_left, function.comp_apply], + refine congr_arg u (fin.ext _), + simp }, + { exact (fin.cast hv r).elim0' } +end + +@[simp] lemma append_elim0' {α : Type*} (u : fin m → α) : + append u fin.elim0' = u ∘ fin.cast (add_zero _) := +append_right_nil _ _ rfl + +lemma append_left_nil {α : Type*} (u : fin m → α) (v : fin n → α) (hu : m = 0) : + append u v = v ∘ fin.cast (by rw [hu, zero_add]) := +begin + refine funext (fin.add_cases (λ l, _) (λ r, _)), + { exact (fin.cast hu l).elim0' }, + { rw [append_right, function.comp_apply], + refine congr_arg v (fin.ext _), + simp [hu] }, +end + +@[simp] lemma elim0'_append {α : Type*} (v : fin n → α) : + append fin.elim0' v = v ∘ fin.cast (zero_add _) := +append_left_nil _ _ rfl + +lemma append_assoc {p : ℕ} {α : Type*} (a : fin m → α) (b : fin n → α) (c : fin p → α) : + append (append a b) c = append a (append b c) ∘ fin.cast (add_assoc _ _ _) := +begin + ext i, + rw function.comp_apply, + refine fin.add_cases (λ l, _) (λ r, _) i, + { rw append_left, + refine fin.add_cases (λ ll, _) (λ lr, _) l, + { rw append_left, + simp [cast_add_cast_add] }, + { rw append_right, + simp [cast_add_nat_add], }, }, + { rw append_right, + simp [←nat_add_nat_add] }, +end + +end append + +section repeat + +/-- Repeat `a` `m` times. For example `fin.repeat 2 ![0, 3, 7] = ![0, 3, 7, 0, 3, 7]`. -/ +@[simp] def repeat {α : Type*} (m : ℕ) (a : fin n → α) : fin (m * n) → α +| i := a i.mod_nat + +@[simp] lemma repeat_zero {α : Type*} (a : fin n → α) : + repeat 0 a = fin.elim0' ∘ cast (zero_mul _) := +funext $ λ x, (cast (zero_mul _) x).elim0' + +@[simp] lemma repeat_one {α : Type*} (a : fin n → α) : + repeat 1 a = a ∘ cast (one_mul _) := +begin + generalize_proofs h, + apply funext, + rw (fin.cast h.symm).surjective.forall, + intro i, + simp [mod_nat, nat.mod_eq_of_lt i.is_lt], +end + +lemma repeat_succ {α : Type*} (a : fin n → α) (m : ℕ) : + repeat m.succ a = append a (repeat m a) ∘ cast ((nat.succ_mul _ _).trans (add_comm _ _)) := +begin + generalize_proofs h, + apply funext, + rw (fin.cast h.symm).surjective.forall, + refine fin.add_cases (λ l, _) (λ r, _), + { simp [mod_nat, nat.mod_eq_of_lt l.is_lt], }, + { simp [mod_nat] } +end + +@[simp] lemma repeat_add {α : Type*} (a : fin n → α) (m₁ m₂ : ℕ) : + repeat (m₁ + m₂) a = append (repeat m₁ a) (repeat m₂ a) ∘ cast (add_mul _ _ _) := +begin + generalize_proofs h, + apply funext, + rw (fin.cast h.symm).surjective.forall, + refine fin.add_cases (λ l, _) (λ r, _), + { simp [mod_nat, nat.mod_eq_of_lt l.is_lt], }, + { simp [mod_nat, nat.add_mod] } +end + +end repeat end tuple diff --git a/src/data/fin/vec_notation.lean b/src/data/fin/vec_notation.lean index e19013da44475..dbc377afeeb7e 100644 --- a/src/data/fin/vec_notation.lean +++ b/src/data/fin/vec_notation.lean @@ -45,8 +45,7 @@ variables {α : Type u} section matrix_notation /-- `![]` is the vector with no entries. -/ -def vec_empty : fin 0 → α := -fin_zero_elim +def vec_empty : fin 0 → α := fin.elim0' /-- `vec_cons h t` prepends an entry `h` to a vector `t`. @@ -174,16 +173,43 @@ of elements by virtue of the semantics of `bit0` and `bit1` and of addition on `fin n`). -/ -@[simp] lemma empty_append (v : fin n → α) : fin.append (zero_add _).symm ![] v = v := -by { ext, simp [fin.append] } +/-- `vec_append ho u v` appends two vectors of lengths `m` and `n` to produce +one of length `o = m + n`. This is a variant of `fin.append` with an additional `ho` argument, +which provides control of definitional equality for the vector length. + +This turns out to be helpful when providing simp lemmas to reduce `![a, b, c] n`, and also means +that `vec_append ho u v 0` is valid. `fin.append u v 0` is not valid in this case because there is +no `has_zero (fin (m + n))` instance. -/ +def vec_append {α : Type*} {o : ℕ} (ho : o = m + n) (u : fin m → α) (v : fin n → α) : fin o → α := +fin.append u v ∘ fin.cast ho + +lemma vec_append_eq_ite {α : Type*} {o : ℕ} (ho : o = m + n) (u : fin m → α) (v : fin n → α) : + vec_append ho u v = λ i, + if h : (i : ℕ) < m + then u ⟨i, h⟩ + else v ⟨(i : ℕ) - m, (tsub_lt_iff_left (le_of_not_lt h)).2 (ho ▸ i.property)⟩ := +begin + ext i, + rw [vec_append, fin.append, function.comp_apply, fin.add_cases], + congr' with hi, + simp only [eq_rec_constant], + refl, +end + +@[simp] lemma vec_append_apply_zero {α : Type*} {o : ℕ} (ho : (o + 1) = (m + 1) + n) + (u : fin (m + 1) → α) (v : fin n → α) : + vec_append ho u v 0 = u 0 := rfl + +@[simp] lemma empty_vec_append (v : fin n → α) : vec_append (zero_add _).symm ![] v = v := +by { ext, simp [vec_append_eq_ite] } -@[simp] lemma cons_append (ho : o + 1 = m + 1 + n) (x : α) (u : fin m → α) (v : fin n → α) : - fin.append ho (vec_cons x u) v = - vec_cons x (fin.append (by rwa [add_assoc, add_comm 1, ←add_assoc, +@[simp] lemma cons_vec_append (ho : o + 1 = m + 1 + n) (x : α) (u : fin m → α) (v : fin n → α) : + vec_append ho (vec_cons x u) v = + vec_cons x (vec_append (by rwa [add_assoc, add_comm 1, ←add_assoc, add_right_cancel_iff] at ho) u v) := begin ext i, - simp_rw [fin.append], + simp_rw [vec_append_eq_ite], split_ifs with h, { rcases i with ⟨⟨⟩ | i, hi⟩, { simp }, @@ -205,10 +231,10 @@ only alternate elements (odd-numbered). -/ def vec_alt1 (hm : m = n + n) (v : fin m → α) (k : fin n) : α := v ⟨(k : ℕ) + k + 1, hm.symm ▸ nat.add_succ_lt_add k.property k.property⟩ -lemma vec_alt0_append (v : fin n → α) : vec_alt0 rfl (fin.append rfl v v) = v ∘ bit0 := +lemma vec_alt0_vec_append (v : fin n → α) : vec_alt0 rfl (vec_append rfl v v) = v ∘ bit0 := begin ext i, - simp_rw [function.comp, bit0, vec_alt0, fin.append], + simp_rw [function.comp, bit0, vec_alt0, vec_append_eq_ite], split_ifs with h; congr, { rw fin.coe_mk at h, simp only [fin.ext_iff, fin.coe_add, fin.coe_mk], @@ -220,10 +246,10 @@ begin exact add_lt_add i.property i.property } end -lemma vec_alt1_append (v : fin (n + 1) → α) : vec_alt1 rfl (fin.append rfl v v) = v ∘ bit1 := +lemma vec_alt1_vec_append (v : fin (n + 1) → α) : vec_alt1 rfl (vec_append rfl v v) = v ∘ bit1 := begin ext i, - simp_rw [function.comp, vec_alt1, fin.append], + simp_rw [function.comp, vec_alt1, vec_append_eq_ite], cases n, { simp, congr }, { split_ifs with h; simp_rw [bit1, bit0]; congr, @@ -248,12 +274,12 @@ end by simp [vec_head, vec_alt1] @[simp] lemma cons_vec_bit0_eq_alt0 (x : α) (u : fin n → α) (i : fin (n + 1)) : - vec_cons x u (bit0 i) = vec_alt0 rfl (fin.append rfl (vec_cons x u) (vec_cons x u)) i := -by rw vec_alt0_append + vec_cons x u (bit0 i) = vec_alt0 rfl (vec_append rfl (vec_cons x u) (vec_cons x u)) i := +by rw vec_alt0_vec_append @[simp] lemma cons_vec_bit1_eq_alt1 (x : α) (u : fin n → α) (i : fin (n + 1)) : - vec_cons x u (bit1 i) = vec_alt1 rfl (fin.append rfl (vec_cons x u) (vec_cons x u)) i := -by rw vec_alt1_append + vec_cons x u (bit1 i) = vec_alt1 rfl (vec_append rfl (vec_cons x u) (vec_cons x u)) i := +by rw vec_alt1_vec_append @[simp] lemma cons_vec_alt0 (h : m + 1 + 1 = (n + 1) + (n + 1)) (x y : α) (u : fin m → α) : vec_alt0 h (vec_cons x (vec_cons y u)) = vec_cons x (vec_alt0 diff --git a/src/linear_algebra/cross_product.lean b/src/linear_algebra/cross_product.lean index 91eac76bb2ce3..67f47a6f8ec05 100644 --- a/src/linear_algebra/cross_product.lean +++ b/src/linear_algebra/cross_product.lean @@ -96,8 +96,8 @@ lemma triple_product_permutation (u v w : fin 3 → R) : u ⬝ᵥ (v ×₃ w) = v ⬝ᵥ (w ×₃ u) := begin simp only [cross_apply, vec3_dot_product, - matrix.head_cons, matrix.cons_vec_bit0_eq_alt0, matrix.empty_append, matrix.cons_val_one, - matrix.cons_vec_alt0, matrix.cons_append, matrix.cons_val_zero], + matrix.head_cons, matrix.cons_vec_bit0_eq_alt0, matrix.empty_vec_append, matrix.cons_val_one, + matrix.cons_vec_alt0, matrix.cons_vec_append, matrix.cons_val_zero], ring, end @@ -108,8 +108,8 @@ theorem triple_product_eq_det (u v w : fin 3 → R) : begin simp only [vec3_dot_product, cross_apply, matrix.det_fin_three, matrix.head_cons, matrix.cons_vec_bit0_eq_alt0, matrix.empty_vec_alt0, matrix.cons_vec_alt0, - matrix.vec_head_vec_alt0, fin.fin_append_apply_zero, matrix.empty_append, matrix.cons_append, - matrix.cons_val', matrix.cons_val_one, matrix.cons_val_zero], + matrix.vec_head_vec_alt0, matrix.vec_append_apply_zero, matrix.empty_vec_append, + matrix.cons_vec_append, matrix.cons_val', matrix.cons_val_one, matrix.cons_val_zero], ring, end @@ -117,8 +117,8 @@ end theorem cross_dot_cross (u v w x : fin 3 → R) : (u ×₃ v) ⬝ᵥ (w ×₃ x) = (u ⬝ᵥ w) * (v ⬝ᵥ x) - (u ⬝ᵥ x) * (v ⬝ᵥ w) := begin - simp only [vec3_dot_product, cross_apply, cons_append, cons_vec_bit0_eq_alt0, - cons_val_one, cons_vec_alt0, linear_map.mk₂_apply, cons_val_zero, head_cons, empty_append], + simp only [vec3_dot_product, cross_apply, cons_vec_append, cons_vec_bit0_eq_alt0, + cons_val_one, cons_vec_alt0, linear_map.mk₂_apply, cons_val_zero, head_cons, empty_vec_append], ring_nf, end diff --git a/src/model_theory/encoding.lean b/src/model_theory/encoding.lean index 6b7c3527220d9..b23894d20b87e 100644 --- a/src/model_theory/encoding.lean +++ b/src/model_theory/encoding.lean @@ -74,7 +74,7 @@ begin (fin_range n).map (option.some ∘ ts) ++ list_decode l, { induction (fin_range n) with i l' l'ih, { refl }, - { rw [cons_bind, append_assoc, ih, map_cons, l'ih, cons_append] } }, + { rw [cons_bind, list.append_assoc, ih, map_cons, l'ih, cons_append] } }, have h' : ∀ i, (list_decode ((fin_range n).bind (λ (i : fin n), (ts i).list_encode) ++ l)).nth ↑i = some (some (ts i)), { intro i, @@ -268,7 +268,7 @@ begin rw [list.drop_append_eq_append_drop, length_map, length_fin_range, nat.sub_self, drop, drop_eq_nil_of_le, nil_append], rw [length_map, length_fin_range], }, }, - { rw [list_encode, append_assoc, cons_append, list_decode], + { rw [list_encode, list.append_assoc, cons_append, list_decode], simp only [subtype.val_eq_coe] at *, rw [(ih1 _).1, (ih1 _).2, (ih2 _).1, (ih2 _).2, sigma_imp, dif_pos rfl], exact ⟨rfl, rfl⟩, }, diff --git a/src/number_theory/legendre_symbol/gauss_sum.lean b/src/number_theory/legendre_symbol/gauss_sum.lean index 1045a100f6a18..5a1ae965e601a 100644 --- a/src/number_theory/legendre_symbol/gauss_sum.lean +++ b/src/number_theory/legendre_symbol/gauss_sum.lean @@ -290,7 +290,7 @@ begin { ext, congr, apply pow_one }, convert_to (0 + 1 * τ ^ 1 + 0 + (-1) * τ ^ 3 + 0 + (-1) * τ ^ 5 + 0 + 1 * τ ^ 7) ^ 2 = _, { simp only [χ₈_apply, matrix.cons_val_zero, matrix.cons_val_one, matrix.head_cons, - matrix.cons_vec_bit0_eq_alt0, matrix.cons_vec_bit1_eq_alt1, matrix.cons_append, + matrix.cons_vec_bit0_eq_alt0, matrix.cons_vec_bit1_eq_alt1, matrix.cons_vec_append, matrix.cons_vec_alt0, matrix.cons_vec_alt1, int.cast_zero, int.cast_one, int.cast_neg, zero_mul], refl }, convert_to 8 + (τ ^ 4 + 1) * (τ ^ 10 - 2 * τ ^ 8 - 2 * τ ^ 6 + 6 * τ ^ 4 + τ ^ 2 - 8) = _, diff --git a/src/order/jordan_holder.lean b/src/order/jordan_holder.lean index fade2bc333fae..477d33bd61430 100644 --- a/src/order/jordan_holder.lean +++ b/src/order/jordan_holder.lean @@ -411,89 +411,91 @@ begin ext; simp [this] end -lemma append_cast_add_aux - {s₁ s₂ : composition_series X} - (i : fin s₁.length) : - fin.append (nat.add_succ _ _).symm (s₁ ∘ fin.cast_succ) s₂ - (fin.cast_add s₂.length i).cast_succ = s₁ i.cast_succ := -by { cases i, simp [fin.append, *] } +section fin_lemmas +-- TODO: move these to `vec_notation` and rename them to better describe their statement -lemma append_succ_cast_add_aux - {s₁ s₂ : composition_series X} - (i : fin s₁.length) - (h : s₁ (fin.last _) = s₂ 0) : - fin.append (nat.add_succ _ _).symm (s₁ ∘ fin.cast_succ) s₂ - (fin.cast_add s₂.length i).succ = s₁ i.succ := +variables {α : Type*} {m n : ℕ} (a : fin m.succ → α) (b : fin n.succ → α) + +lemma append_cast_add_aux (i : fin m) : + matrix.vec_append (nat.add_succ _ _).symm (a ∘ fin.cast_succ) b + (fin.cast_add n i).cast_succ = a i.cast_succ := +by { cases i, simp [matrix.vec_append_eq_ite, *] } + +lemma append_succ_cast_add_aux (i : fin m) (h : a (fin.last _) = b 0) : + matrix.vec_append (nat.add_succ _ _).symm (a ∘ fin.cast_succ) b + (fin.cast_add n i).succ = a i.succ := begin cases i with i hi, - simp only [fin.append, hi, fin.succ_mk, function.comp_app, fin.cast_succ_mk, + simp only [matrix.vec_append_eq_ite, hi, fin.succ_mk, function.comp_app, fin.cast_succ_mk, fin.coe_mk, fin.cast_add_mk], split_ifs, { refl }, - { have : i + 1 = s₁.length, from le_antisymm hi (le_of_not_gt h_1), - calc s₂ ⟨i + 1 - s₁.length, by simp [this]⟩ - = s₂ 0 : congr_arg s₂ (by simp [fin.ext_iff, this]) - ... = s₁ (fin.last _) : h.symm - ... = _ : congr_arg s₁ (by simp [fin.ext_iff, this]) } + { have : i + 1 = m, from le_antisymm hi (le_of_not_gt h_1), + calc b ⟨i + 1 - m, by simp [this]⟩ + = b 0 : congr_arg b (by simp [fin.ext_iff, this]) + ... = a (fin.last _) : h.symm + ... = _ : congr_arg a (by simp [fin.ext_iff, this]) } end -lemma append_nat_add_aux - {s₁ s₂ : composition_series X} - (i : fin s₂.length) : - fin.append (nat.add_succ _ _).symm (s₁ ∘ fin.cast_succ) s₂ - (fin.nat_add s₁.length i).cast_succ = s₂ i.cast_succ := +lemma append_nat_add_aux (i : fin n) : + matrix.vec_append (nat.add_succ _ _).symm (a ∘ fin.cast_succ) b + (fin.nat_add m i).cast_succ = b i.cast_succ := begin cases i, - simp only [fin.append, nat.not_lt_zero, fin.nat_add_mk, add_lt_iff_neg_left, + simp only [matrix.vec_append_eq_ite, nat.not_lt_zero, fin.nat_add_mk, add_lt_iff_neg_left, add_tsub_cancel_left, dif_neg, fin.cast_succ_mk, not_false_iff, fin.coe_mk] end -lemma append_succ_nat_add_aux - {s₁ s₂ : composition_series X} - (i : fin s₂.length) : - fin.append (nat.add_succ _ _).symm (s₁ ∘ fin.cast_succ) s₂ - (fin.nat_add s₁.length i).succ = s₂ i.succ := +lemma append_succ_nat_add_aux (i : fin n) : + matrix.vec_append (nat.add_succ _ _).symm (a ∘ fin.cast_succ) b + (fin.nat_add m i).succ = b i.succ := begin cases i with i hi, - simp only [fin.append, add_assoc, nat.not_lt_zero, fin.nat_add_mk, add_lt_iff_neg_left, - add_tsub_cancel_left, fin.succ_mk, dif_neg, not_false_iff, fin.coe_mk] + simp only [matrix.vec_append_eq_ite, add_assoc, nat.not_lt_zero, fin.nat_add_mk, + add_lt_iff_neg_left, add_tsub_cancel_left, fin.succ_mk, dif_neg, not_false_iff, fin.coe_mk] end +end fin_lemmas + /-- Append two composition series `s₁` and `s₂` such that the least element of `s₁` is the maximum element of `s₂`. -/ @[simps length] def append (s₁ s₂ : composition_series X) (h : s₁.top = s₂.bot) : composition_series X := { length := s₁.length + s₂.length, - series := fin.append (nat.add_succ _ _).symm (s₁ ∘ fin.cast_succ) s₂, + series := matrix.vec_append (nat.add_succ _ _).symm (s₁ ∘ fin.cast_succ) s₂, step' := λ i, begin refine fin.add_cases _ _ i, { intro i, - rw [append_succ_cast_add_aux _ h, append_cast_add_aux], + rw [append_succ_cast_add_aux _ _ _ h, append_cast_add_aux], exact s₁.step i }, { intro i, rw [append_nat_add_aux, append_succ_nat_add_aux], exact s₂.step i } end } +lemma coe_append (s₁ s₂ : composition_series X) (h) : + ⇑(s₁.append s₂ h) = matrix.vec_append (nat.add_succ _ _).symm (s₁ ∘ fin.cast_succ) s₂ := +rfl + @[simp] lemma append_cast_add {s₁ s₂ : composition_series X} (h : s₁.top = s₂.bot) (i : fin s₁.length) : append s₁ s₂ h (fin.cast_add s₂.length i).cast_succ = s₁ i.cast_succ := -append_cast_add_aux i +by rw [coe_append, append_cast_add_aux _ _ i] @[simp] lemma append_succ_cast_add {s₁ s₂ : composition_series X} (h : s₁.top = s₂.bot) (i : fin s₁.length) : append s₁ s₂ h (fin.cast_add s₂.length i).succ = s₁ i.succ := -append_succ_cast_add_aux i h +by rw [coe_append, append_succ_cast_add_aux _ _ _ h] @[simp] lemma append_nat_add {s₁ s₂ : composition_series X} (h : s₁.top = s₂.bot) (i : fin s₂.length) : append s₁ s₂ h (fin.nat_add s₁.length i).cast_succ = s₂ i.cast_succ := -append_nat_add_aux i +by rw [coe_append, append_nat_add_aux _ _ i] @[simp] lemma append_succ_nat_add {s₁ s₂ : composition_series X} (h : s₁.top = s₂.bot) (i : fin s₂.length) : append s₁ s₂ h (fin.nat_add s₁.length i).succ = s₂ i.succ := -append_succ_nat_add_aux i +by rw [coe_append, append_succ_nat_add_aux _ _ i] /-- Add an element to the top of a `composition_series` -/ @[simps length] def snoc (s : composition_series X) (x : X) From 56ca12ccf538f9e95bc2b64540ebac99e2c7606d Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 16 Jan 2023 20:04:37 +0000 Subject: [PATCH 0261/1291] =?UTF-8?q?feat(group=5Ftheory/perm/cycle/basic)?= =?UTF-8?q?:=20Partition=20`s=20=C3=97=20s`=20into=20shifted=20diagonals?= =?UTF-8?q?=20(#17910)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Given a cycle `f` on a finset `s`, we can partition the square `s ×ˢ s` into `(λ a, (a, (f ^ n) a)) '' s`, ranging over `n < card s`. This translates to an expansion of `∑ i in s, f i * ∑ i in s, g i` which is key to Chebyshev's sum inequality (#13187). For a finset with five elements, the picture looks like this: ``` 01234 40123 34012 23401 12340 ``` --- src/group_theory/perm/cycle/basic.lean | 86 +++++++++++++++++++++++++- 1 file changed, 85 insertions(+), 1 deletion(-) diff --git a/src/group_theory/perm/cycle/basic.lean b/src/group_theory/perm/cycle/basic.lean index caa1530b9dcef..29a35fa756831 100644 --- a/src/group_theory/perm/cycle/basic.lean +++ b/src/group_theory/perm/cycle/basic.lean @@ -3,6 +3,7 @@ Copyright (c) 2019 Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes, Yaël Dillies -/ +import algebra.module.big_operators import data.finset.noncomm_prod import data.fintype.perm import data.int.modeq @@ -49,8 +50,9 @@ The following two definitions require that `β` is a `fintype`: -/ open equiv function finset +open_locale big_operators -variables {α β : Type*} +variables {ι α β : Type*} namespace equiv.perm @@ -1609,3 +1611,85 @@ end end fixed_points end equiv.perm + +open equiv + +namespace set +variables {f : perm α} {s : set α} + +lemma prod_self_eq_Union_perm (hf : f.is_cycle_on s) : + s ×ˢ s = ⋃ n : ℤ, (λ a, (a, (f ^ n) a)) '' s := +begin + ext ⟨a, b⟩, + simp only [mem_prod, mem_Union, mem_image], + refine ⟨λ hx, _, _⟩, + { obtain ⟨n, rfl⟩ := hf.2 hx.1 hx.2, + exact ⟨_, _, hx.1, rfl⟩ }, + { rintro ⟨n, a, ha, ⟨⟩⟩, + exact ⟨ha, (hf.1.perm_zpow _).maps_to ha⟩ } +end + +end set + +namespace finset +variables {f : perm α} {s : finset α} + +lemma product_self_eq_disj_Union_perm_aux (hf : f.is_cycle_on s) : + (range s.card : set ℕ).pairwise_disjoint + (λ k, s.map ⟨λ i, (i, (f ^ k) i), λ i j, congr_arg prod.fst⟩) := +begin + obtain hs | hs := (s : set α).subsingleton_or_nontrivial, + { refine set.subsingleton.pairwise _ _, + simp_rw [set.subsingleton, mem_coe, ←card_le_one] at ⊢ hs, + rwa card_range }, + classical, + rintro m hm n hn hmn, + simp only [disjoint_left, function.on_fun, mem_map, function.embedding.coe_fn_mk, exists_prop, + not_exists, not_and, forall_exists_index, and_imp, prod.forall, prod.mk.inj_iff], + rintro _ _ _ - rfl rfl a ha rfl h, + rw [hf.pow_apply_eq_pow_apply ha] at h, + rw [mem_coe, mem_range] at hm hn, + exact hmn.symm (h.eq_of_lt_of_lt hn hm), +end + +/-- +We can partition the square `s ×ˢ s` into shifted diagonals as such: +``` +01234 +40123 +34012 +23401 +12340 +``` + +The diagonals are given by the cycle `f`. +-/ +lemma product_self_eq_disj_Union_perm (hf : f.is_cycle_on s) : + s ×ˢ s = + (range s.card).disj_Union (λ k, s.map ⟨λ i, (i, (f ^ k) i), λ i j, congr_arg prod.fst⟩) + (product_self_eq_disj_Union_perm_aux hf) := +begin + ext ⟨a, b⟩, + simp only [mem_product, equiv.perm.coe_pow, mem_disj_Union, mem_range, mem_map, + function.embedding.coe_fn_mk, prod.mk.inj_iff, exists_prop], + refine ⟨λ hx, _, _⟩, + { obtain ⟨n, hn, rfl⟩ := hf.exists_pow_eq hx.1 hx.2, + exact ⟨n, hn, a, hx.1, rfl, by rw f.iterate_eq_pow⟩ }, + { rintro ⟨n, -, a, ha, rfl, rfl⟩, + exact ⟨ha, (hf.1.iterate _).maps_to ha⟩ } +end + +end finset + +namespace finset +variables [semiring α] [add_comm_monoid β] [module α β] {s : finset ι} {σ : perm ι} + +lemma sum_smul_sum_eq_sum_perm (hσ : σ.is_cycle_on s) (f : ι → α) (g : ι → β) : + (∑ i in s, f i) • ∑ i in s, g i = ∑ k in range s.card, ∑ i in s, f i • g ((σ ^ k) i) := +by { simp_rw [sum_smul_sum, product_self_eq_disj_Union_perm hσ, sum_disj_Union, sum_map], refl } + +lemma sum_mul_sum_eq_sum_perm (hσ : σ.is_cycle_on s) (f g : ι → α) : + (∑ i in s, f i) * ∑ i in s, g i = ∑ k in range s.card, ∑ i in s, f i * g ((σ ^ k) i) := +sum_smul_sum_eq_sum_perm hσ f g + +end finset From 1447cae870f372074e480de1acbeb51de0077698 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Mon, 16 Jan 2023 20:04:40 +0000 Subject: [PATCH 0262/1291] chore(data/list): reorder arguments of `list.nthd` (#18182) Sync it with `List.getD` in Mathlib 4. --- src/computability/primrec.lean | 9 +++------ src/data/list/basic.lean | 24 ++++++++++++------------ src/data/list/defs.lean | 10 +++++----- 3 files changed, 20 insertions(+), 23 deletions(-) diff --git a/src/computability/primrec.lean b/src/computability/primrec.lean index 65a86d454f31d..669d748677efa 100644 --- a/src/computability/primrec.lean +++ b/src/computability/primrec.lean @@ -940,13 +940,10 @@ this.to₂.of_eq $ λ l n, begin { apply IH } end -theorem list_nthd (d : α) : primrec₂ (list.nthd d) := +theorem list_nthd (d : α) : primrec₂ (λ l n, list.nthd l n d) := begin - suffices : list.nthd d = λ l n, (list.nth l n).get_or_else d, - { rw this, - exact option_get_or_else.comp₂ list_nth (const _) }, - funext, - exact list.nthd_eq_get_or_else_nth _ _ _ + simp only [list.nthd_eq_get_or_else_nth], + exact option_get_or_else.comp₂ list_nth (const _) end theorem list_inth [inhabited α] : primrec₂ (@list.inth α _) := diff --git a/src/data/list/basic.lean b/src/data/list/basic.lean index 61268e5fc965b..242db1cf4034d 100644 --- a/src/data/list/basic.lean +++ b/src/data/list/basic.lean @@ -4103,13 +4103,13 @@ section nthd variables (l : list α) (x : α) (xs : list α) (d : α) (n : ℕ) -@[simp] lemma nthd_nil : nthd d [] n = d := rfl +@[simp] lemma nthd_nil : nthd [] n d = d := rfl -@[simp] lemma nthd_cons_zero : nthd d (x::xs) 0 = x := rfl +@[simp] lemma nthd_cons_zero : nthd (x::xs) 0 d = x := rfl -@[simp] lemma nthd_cons_succ : nthd d (x::xs) (n + 1) = nthd d xs n := rfl +@[simp] lemma nthd_cons_succ : nthd (x::xs) (n + 1) d = nthd xs n d := rfl -lemma nthd_eq_nth_le {n : ℕ} (hn : n < l.length) : l.nthd d n = l.nth_le n hn := +lemma nthd_eq_nth_le {n : ℕ} (hn : n < l.length) : l.nthd n d = l.nth_le n hn := begin induction l with hd tl IH generalizing n, { exact absurd hn (not_lt_of_ge (nat.zero_le _)) }, @@ -4118,7 +4118,7 @@ begin { exact IH _ } } end -lemma nthd_eq_default {n : ℕ} (hn : l.length ≤ n) : l.nthd d n = d := +lemma nthd_eq_default {n : ℕ} (hn : l.length ≤ n) : l.nthd n d = d := begin induction l with hd tl IH generalizing n, { exact nthd_nil _ _ }, @@ -4130,12 +4130,12 @@ end /-- An empty list can always be decidably checked for the presence of an element. Not an instance because it would clash with `decidable_eq α`. -/ def decidable_nthd_nil_ne {α} (a : α) : decidable_pred - (λ (i : ℕ), nthd a ([] : list α) i ≠ a) := λ i, is_false $ λ H, H (nthd_nil _ _) + (λ (i : ℕ), nthd ([] : list α) i a ≠ a) := λ i, is_false $ λ H, H (nthd_nil _ _) -@[simp] lemma nthd_singleton_default_eq (n : ℕ) : [d].nthd d n = d := +@[simp] lemma nthd_singleton_default_eq (n : ℕ) : [d].nthd n d = d := by { cases n; simp } -@[simp] lemma nthd_replicate_default_eq (r n : ℕ) : (replicate r d).nthd d n = d := +@[simp] lemma nthd_replicate_default_eq (r n : ℕ) : (replicate r d).nthd n d = d := begin induction r with r IH generalizing n, { simp }, @@ -4145,11 +4145,11 @@ end lemma nthd_append (l l' : list α) (d : α) (n : ℕ) (h : n < l.length) (h' : n < (l ++ l').length := h.trans_le ((length_append l l').symm ▸ le_self_add)) : - (l ++ l').nthd d n = l.nthd d n := + (l ++ l').nthd n d = l.nthd n d := by rw [nthd_eq_nth_le _ _ h', nth_le_append h' h, nthd_eq_nth_le] lemma nthd_append_right (l l' : list α) (d : α) (n : ℕ) (h : l.length ≤ n) : - (l ++ l').nthd d n = l'.nthd d (n - l.length) := + (l ++ l').nthd n d = l'.nthd (n - l.length) d := begin cases lt_or_le _ _ with h' h', { rw [nthd_eq_nth_le _ _ h', nth_le_append_right h h', nthd_eq_nth_le] }, @@ -4158,7 +4158,7 @@ begin end lemma nthd_eq_get_or_else_nth (n : ℕ) : - l.nthd d n = (l.nth n).get_or_else d := + l.nthd n d = (l.nth n).get_or_else d := begin cases lt_or_le _ _ with h h, { rw [nthd_eq_nth_le _ _ h, nth_le_nth h, option.get_or_else_some] }, @@ -4181,7 +4181,7 @@ lemma inth_eq_nth_le {n : ℕ} (hn : n < l.length) : l.inth n = l.nth_le n hn := lemma inth_eq_default {n : ℕ} (hn : l.length ≤ n) : l.inth n = default := nthd_eq_default _ _ hn -lemma nthd_default_eq_inth : l.nthd default = l.inth := rfl +lemma nthd_default_eq_inth : l.nthd n default = l.inth n := rfl lemma inth_append (l l' : list α) (n : ℕ) (h : n < l.length) (h' : n < (l ++ l').length := h.trans_le ((length_append l l').symm ▸ le_self_add)) : diff --git a/src/data/list/defs.lean b/src/data/list/defs.lean index d548a377c57b7..e053c0de9a350 100644 --- a/src/data/list/defs.lean +++ b/src/data/list/defs.lean @@ -76,14 +76,14 @@ def to_array (l : list α) : array l.length α := /-- "default" `nth` function: returns `d` instead of `none` in the case that the index is out of bounds. -/ -def nthd (d : α) : Π (l : list α) (n : ℕ), α -| [] _ := d -| (x::xs) 0 := x -| (x::xs) (n + 1) := nthd xs n +def nthd : Π (l : list α) (n : ℕ) (d : α), α +| [] _ d := d +| (x::xs) 0 d := x +| (x::xs) (n + 1) d := nthd xs n d /-- "inhabited" `nth` function: returns `default` instead of `none` in the case that the index is out of bounds. -/ -def inth [h : inhabited α] (l : list α) (n : nat) : α := nthd default l n +def inth [h : inhabited α] (l : list α) (n : nat) : α := nthd l n default /-- Apply a function to the nth tail of `l`. Returns the input without using `f` if the index is larger than the length of the list. From e6c26fe38ddd3d4175caf4bd24d685901eb85fe5 Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Mon, 16 Jan 2023 23:40:58 +0000 Subject: [PATCH 0263/1291] chore(analysis/convex/body): downgrade imports (#18185) Downgrade imports, by generalizing the statements from a normed space to a topological vector space. --- src/analysis/convex/body.lean | 12 ++++++------ 1 file changed, 6 insertions(+), 6 deletions(-) diff --git a/src/analysis/convex/body.lean b/src/analysis/convex/body.lean index 7fa9402b26c89..ec6b5060befa9 100644 --- a/src/analysis/convex/body.lean +++ b/src/analysis/convex/body.lean @@ -4,17 +4,16 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Paul A. Reichert -/ import analysis.convex.basic -import analysis.normed_space.basic import data.real.nnreal -import data.set.pointwise.basic -import topology.subset_properties +import topology.algebra.module.basic +import topology.instances.real /-! # convex bodies This file contains the definition of the type `convex_body V` consisting of -convex, compact, nonempty subsets of a real normed space `V`. +convex, compact, nonempty subsets of a real topological vector space `V`. `convex_body V` is a module over the nonnegative reals (`nnreal`). @@ -31,10 +30,11 @@ convex, convex body open_locale pointwise open_locale nnreal -variables (V : Type*) [seminormed_add_comm_group V] [normed_space ℝ V] +variables (V : Type*) [topological_space V] [add_comm_group V] [has_continuous_add V] + [module ℝ V] [has_continuous_smul ℝ V] /-- -Let `V` be a normed space. A subset of `V` is a convex body if and only if +Let `V` be a real topological vector space. A subset of `V` is a convex body if and only if it is convex, compact, and nonempty. -/ structure convex_body := From 4b99fe0a1096dc52abe68e65107220e604ea49b2 Mon Sep 17 00:00:00 2001 From: JovanGerb Date: Mon, 16 Jan 2023 23:40:59 +0000 Subject: [PATCH 0264/1291] chore(topology/metric_space/isometry): rename isometric to isometry_equiv (#18191) The name for isometric bijections is changed to `isometry_equiv`, to be consistent with `linear_isometry_equiv` and `affine_isometry_equiv`. --- src/analysis/normed/group/add_torsor.lean | 14 +++---- src/analysis/normed/group/basic.lean | 29 +++++++------- src/analysis/normed_space/add_torsor.lean | 2 +- .../normed_space/affine_isometry.lean | 14 +++---- .../normed_space/linear_isometry.lean | 38 +++++++++---------- src/analysis/normed_space/mazur_ulam.lean | 28 +++++++------- .../conditional_expectation/basic.lean | 4 +- src/measure_theory/measure/hausdorff.lean | 10 ++--- src/topology/continuous_function/compact.lean | 11 +++--- .../metric_space/gromov_hausdorff.lean | 28 +++++++------- .../metric_space/hausdorff_dimension.lean | 4 +- src/topology/metric_space/isometry.lean | 18 ++++----- 12 files changed, 101 insertions(+), 99 deletions(-) diff --git a/src/analysis/normed/group/add_torsor.lean b/src/analysis/normed/group/add_torsor.lean index 0e1f3e2a565d9..527f55a44b4ff 100644 --- a/src/analysis/normed/group/add_torsor.lean +++ b/src/analysis/normed/group/add_torsor.lean @@ -85,7 +85,7 @@ by rw [dist_comm, dist_vadd_left] /-- Isometry between the tangent space `V` of a (semi)normed add torsor `P` and `P` given by addition/subtraction of `x : P`. -/ -@[simps] def isometric.vadd_const (x : P) : V ≃ᵢ P := +@[simps] def isometry_equiv.vadd_const (x : P) : V ≃ᵢ P := { to_equiv := equiv.vadd_const x, isometry_to_fun := isometry.of_dist_eq $ λ _ _, dist_vadd_cancel_right _ _ _ } @@ -94,7 +94,7 @@ section variable (P) /-- Self-isometry of a (semi)normed add torsor given by addition of a constant vector `x`. -/ -@[simps] def isometric.const_vadd (x : V) : P ≃ᵢ P := +@[simps] def isometry_equiv.const_vadd (x : V) : P ≃ᵢ P := { to_equiv := equiv.const_vadd P x, isometry_to_fun := isometry.of_dist_eq $ λ _ _, dist_vadd_cancel_left _ _ _ } @@ -105,12 +105,12 @@ by rw [dist_eq_norm, vsub_sub_vsub_cancel_left, dist_comm, dist_eq_norm_vsub V] /-- Isometry between the tangent space `V` of a (semi)normed add torsor `P` and `P` given by subtraction from `x : P`. -/ -@[simps] def isometric.const_vsub (x : P) : P ≃ᵢ V := +@[simps] def isometry_equiv.const_vsub (x : P) : P ≃ᵢ V := { to_equiv := equiv.const_vsub x, isometry_to_fun := isometry.of_dist_eq $ λ y z, dist_vsub_cancel_left _ _ _ } @[simp] lemma dist_vsub_cancel_right (x y z : P) : dist (x -ᵥ z) (y -ᵥ z) = dist x y := -(isometric.vadd_const z).symm.dist_eq x y +(isometry_equiv.vadd_const z).symm.dist_eq x y section pointwise @@ -118,15 +118,15 @@ open_locale pointwise @[simp] lemma vadd_ball (x : V) (y : P) (r : ℝ) : x +ᵥ metric.ball y r = metric.ball (x +ᵥ y) r := -(isometric.const_vadd P x).image_ball y r +(isometry_equiv.const_vadd P x).image_ball y r @[simp] lemma vadd_closed_ball (x : V) (y : P) (r : ℝ) : x +ᵥ metric.closed_ball y r = metric.closed_ball (x +ᵥ y) r := -(isometric.const_vadd P x).image_closed_ball y r +(isometry_equiv.const_vadd P x).image_closed_ball y r @[simp] lemma vadd_sphere (x : V) (y : P) (r : ℝ) : x +ᵥ metric.sphere y r = metric.sphere (x +ᵥ y) r := -(isometric.const_vadd P x).image_sphere y r +(isometry_equiv.const_vadd P x).image_sphere y r end pointwise diff --git a/src/analysis/normed/group/basic.lean b/src/analysis/normed/group/basic.lean index 3eba78c8551d9..f82dcde4fcc56 100644 --- a/src/analysis/normed/group/basic.lean +++ b/src/analysis/normed/group/basic.lean @@ -493,7 +493,7 @@ def norm_group_seminorm : group_seminorm E := ⟨norm, norm_one', norm_mul_le', variables {E} -namespace isometric +namespace isometry_equiv -- TODO This material is superseded by similar constructions such as -- `affine_isometry_equiv.const_vadd`; deduplicate @@ -504,18 +504,19 @@ protected def mul_right (x : E) : E ≃ᵢ E := .. equiv.mul_right x } @[simp, to_additive] -lemma mul_right_to_equiv (x : E) : (isometric.mul_right x).to_equiv = equiv.mul_right x := rfl +lemma mul_right_to_equiv (x : E) : (isometry_equiv.mul_right x).to_equiv = equiv.mul_right x := rfl @[simp, to_additive] -lemma coe_mul_right (x : E) : (isometric.mul_right x : E → E) = λ y, y * x := rfl +lemma coe_mul_right (x : E) : (isometry_equiv.mul_right x : E → E) = λ y, y * x := rfl -@[to_additive] lemma mul_right_apply (x y : E) : (isometric.mul_right x : E → E) y = y * x := rfl +@[to_additive] lemma mul_right_apply (x y : E) : (isometry_equiv.mul_right x : E → E) y = y * x := +rfl @[simp, to_additive] -lemma mul_right_symm (x : E) : (isometric.mul_right x).symm = isometric.mul_right x⁻¹ := +lemma mul_right_symm (x : E) : (isometry_equiv.mul_right x).symm = isometry_equiv.mul_right x⁻¹ := ext $ λ y, rfl -end isometric +end isometry_equiv @[to_additive] lemma normed_comm_group.tendsto_nhds_one {f : α → E} {l : filter α} : tendsto f l (𝓝 1) ↔ ∀ ε > 0, ∀ᶠ x in l, ‖ f x ‖ < ε := @@ -1107,7 +1108,7 @@ by { ext, simp [mem_closed_ball, set.mem_smul_set, dist_eq_norm_div, div_eq_inv_ by { ext, simp [mem_ball, set.mem_smul_set, dist_eq_norm_div, div_eq_inv_mul, ← eq_inv_mul_iff_mul_eq, mul_assoc], } -namespace isometric +namespace isometry_equiv /-- Multiplication `y ↦ x * y` as an `isometry`. -/ @[to_additive "Addition `y ↦ x + y` as an `isometry`"] @@ -1116,12 +1117,12 @@ protected def mul_left (x : E) : E ≃ᵢ E := to_equiv := equiv.mul_left x } @[simp, to_additive] lemma mul_left_to_equiv (x : E) : - (isometric.mul_left x).to_equiv = equiv.mul_left x := rfl + (isometry_equiv.mul_left x).to_equiv = equiv.mul_left x := rfl -@[simp, to_additive] lemma coe_mul_left (x : E) : ⇑(isometric.mul_left x) = (*) x := rfl +@[simp, to_additive] lemma coe_mul_left (x : E) : ⇑(isometry_equiv.mul_left x) = (*) x := rfl @[simp, to_additive] lemma mul_left_symm (x : E) : - (isometric.mul_left x).symm = isometric.mul_left x⁻¹ := + (isometry_equiv.mul_left x).symm = isometry_equiv.mul_left x⁻¹ := ext $ λ y, rfl variables (E) @@ -1133,11 +1134,11 @@ variables (E) variables {E} -@[simp, to_additive] lemma inv_symm : (isometric.inv E).symm = isometric.inv E := rfl -@[simp, to_additive] lemma inv_to_equiv : (isometric.inv E).to_equiv = equiv.inv E := rfl -@[simp, to_additive] lemma coe_inv : ⇑(isometric.inv E) = has_inv.inv := rfl +@[simp, to_additive] lemma inv_symm : (isometry_equiv.inv E).symm = isometry_equiv.inv E := rfl +@[simp, to_additive] lemma inv_to_equiv : (isometry_equiv.inv E).to_equiv = equiv.inv E := rfl +@[simp, to_additive] lemma coe_inv : ⇑(isometry_equiv.inv E) = has_inv.inv := rfl -end isometric +end isometry_equiv open finset diff --git a/src/analysis/normed_space/add_torsor.lean b/src/analysis/normed_space/add_torsor.lean index 21d02eeb0372b..145e71133e77f 100644 --- a/src/analysis/normed_space/add_torsor.lean +++ b/src/analysis/normed_space/add_torsor.lean @@ -32,7 +32,7 @@ lemma affine_subspace.is_closed_direction_iff (s : affine_subspace 𝕜 Q) : is_closed (s.direction : set W) ↔ is_closed (s : set Q) := begin rcases s.eq_bot_or_nonempty with rfl|⟨x, hx⟩, { simp [is_closed_singleton] }, - rw [← (isometric.vadd_const x).to_homeomorph.symm.is_closed_image, + rw [← (isometry_equiv.vadd_const x).to_homeomorph.symm.is_closed_image, affine_subspace.coe_direction_eq_vsub_set_right hx], refl end diff --git a/src/analysis/normed_space/affine_isometry.lean b/src/analysis/normed_space/affine_isometry.lean index 34389bb722cb6..efb9b75d83be6 100644 --- a/src/analysis/normed_space/affine_isometry.lean +++ b/src/analysis/normed_space/affine_isometry.lean @@ -317,18 +317,18 @@ variables (e : P ≃ᵃⁱ[𝕜] P₂) protected lemma isometry : isometry e := e.to_affine_isometry.isometry -/-- Reinterpret a `affine_isometry_equiv` as an `isometric`. -/ -def to_isometric : P ≃ᵢ P₂ := ⟨e.to_affine_equiv.to_equiv, e.isometry⟩ +/-- Reinterpret a `affine_isometry_equiv` as an `isometry_equiv`. -/ +def to_isometry_equiv : P ≃ᵢ P₂ := ⟨e.to_affine_equiv.to_equiv, e.isometry⟩ -@[simp] lemma coe_to_isometric : ⇑e.to_isometric = e := rfl +@[simp] lemma coe_to_isometry_equiv : ⇑e.to_isometry_equiv = e := rfl include V V₂ lemma range_eq_univ (e : P ≃ᵃⁱ[𝕜] P₂) : set.range e = set.univ := -by { rw ← coe_to_isometric, exact isometric.range_eq_univ _, } +by { rw ← coe_to_isometry_equiv, exact isometry_equiv.range_eq_univ _, } omit V V₂ /-- Reinterpret a `affine_isometry_equiv` as an `homeomorph`. -/ -def to_homeomorph : P ≃ₜ P₂ := e.to_isometric.to_homeomorph +def to_homeomorph : P ≃ₜ P₂ := e.to_isometry_equiv.to_homeomorph @[simp] lemma coe_to_homeomorph : ⇑e.to_homeomorph = e := rfl @@ -351,7 +351,7 @@ instance : inhabited (P ≃ᵃⁱ[𝕜] P) := ⟨refl 𝕜 P⟩ @[simp] lemma coe_refl : ⇑(refl 𝕜 P) = id := rfl @[simp] lemma to_affine_equiv_refl : (refl 𝕜 P).to_affine_equiv = affine_equiv.refl 𝕜 P := rfl -@[simp] lemma to_isometric_refl : (refl 𝕜 P).to_isometric = isometric.refl P := rfl +@[simp] lemma to_isometry_equiv_refl : (refl 𝕜 P).to_isometry_equiv = isometry_equiv.refl P := rfl @[simp] lemma to_homeomorph_refl : (refl 𝕜 P).to_homeomorph = homeomorph.refl P := rfl omit V @@ -365,7 +365,7 @@ def symm : P₂ ≃ᵃⁱ[𝕜] P := @[simp] lemma symm_symm : e.symm.symm = e := ext $ λ x, rfl @[simp] lemma to_affine_equiv_symm : e.to_affine_equiv.symm = e.symm.to_affine_equiv := rfl -@[simp] lemma to_isometric_symm : e.to_isometric.symm = e.symm.to_isometric := rfl +@[simp] lemma to_isometry_equiv_symm : e.to_isometry_equiv.symm = e.symm.to_isometry_equiv := rfl @[simp] lemma to_homeomorph_symm : e.to_homeomorph.symm = e.symm.to_homeomorph := rfl include V₃ diff --git a/src/analysis/normed_space/linear_isometry.lean b/src/analysis/normed_space/linear_isometry.lean index f23d83eb76925..95db39fb82021 100644 --- a/src/analysis/normed_space/linear_isometry.lean +++ b/src/analysis/normed_space/linear_isometry.lean @@ -456,24 +456,24 @@ to_linear_isometry_injective.eq_iff protected lemma isometry : isometry e := e.to_linear_isometry.isometry -/-- Reinterpret a `linear_isometry_equiv` as an `isometric`. -/ -def to_isometric : E ≃ᵢ E₂ := ⟨e.to_linear_equiv.to_equiv, e.isometry⟩ +/-- Reinterpret a `linear_isometry_equiv` as an `isometry_equiv`. -/ +def to_isometry_equiv : E ≃ᵢ E₂ := ⟨e.to_linear_equiv.to_equiv, e.isometry⟩ -lemma to_isometric_injective : - function.injective (to_isometric : (E ≃ₛₗᵢ[σ₁₂] E₂) → E ≃ᵢ E₂) := -λ x y h, coe_injective (congr_arg _ h : ⇑x.to_isometric = _) +lemma to_isometry_equiv_injective : + function.injective (to_isometry_equiv : (E ≃ₛₗᵢ[σ₁₂] E₂) → E ≃ᵢ E₂) := +λ x y h, coe_injective (congr_arg _ h : ⇑x.to_isometry_equiv = _) -@[simp] lemma to_isometric_inj {f g : E ≃ₛₗᵢ[σ₁₂] E₂} : - f.to_isometric = g.to_isometric ↔ f = g := -to_isometric_injective.eq_iff +@[simp] lemma to_isometry_equiv_inj {f g : E ≃ₛₗᵢ[σ₁₂] E₂} : + f.to_isometry_equiv = g.to_isometry_equiv ↔ f = g := +to_isometry_equiv_injective.eq_iff -@[simp] lemma coe_to_isometric : ⇑e.to_isometric = e := rfl +@[simp] lemma coe_to_isometry_equiv : ⇑e.to_isometry_equiv = e := rfl lemma range_eq_univ (e : E ≃ₛₗᵢ[σ₁₂] E₂) : set.range e = set.univ := -by { rw ← coe_to_isometric, exact isometric.range_eq_univ _, } +by { rw ← coe_to_isometry_equiv, exact isometry_equiv.range_eq_univ _, } /-- Reinterpret a `linear_isometry_equiv` as an `homeomorph`. -/ -def to_homeomorph : E ≃ₜ E₂ := e.to_isometric.to_homeomorph +def to_homeomorph : E ≃ₜ E₂ := e.to_isometry_equiv.to_homeomorph lemma to_homeomorph_injective : function.injective (to_homeomorph : (E ≃ₛₗᵢ[σ₁₂] E₂) → E ≃ₜ E₂) := @@ -531,7 +531,7 @@ def symm : E₂ ≃ₛₗᵢ[σ₂₁] E := @[simp] lemma symm_symm : e.symm.symm = e := ext $ λ x, rfl @[simp] lemma to_linear_equiv_symm : e.to_linear_equiv.symm = e.symm.to_linear_equiv := rfl -@[simp] lemma to_isometric_symm : e.to_isometric.symm = e.symm.to_isometric := rfl +@[simp] lemma to_isometry_equiv_symm : e.to_isometry_equiv.symm = e.symm.to_isometry_equiv := rfl @[simp] lemma to_homeomorph_symm : e.to_homeomorph.symm = e.symm.to_homeomorph := rfl /-- See Note [custom simps projection]. We need to specify this projection explicitly in this case, @@ -624,7 +624,7 @@ include σ₂₁ /-- Reinterpret a `linear_isometry_equiv` as a `continuous_linear_equiv`. -/ instance : has_coe_t (E ≃ₛₗᵢ[σ₁₂] E₂) (E ≃SL[σ₁₂] E₂) := -⟨λ e, ⟨e.to_linear_equiv, e.continuous, e.to_isometric.symm.continuous⟩⟩ +⟨λ e, ⟨e.to_linear_equiv, e.continuous, e.to_isometry_equiv.symm.continuous⟩⟩ instance : has_coe_t (E ≃ₛₗᵢ[σ₁₂] E₂) (E →SL[σ₁₂] E₂) := ⟨λ e, ↑(e : E ≃SL[σ₁₂] E₂)⟩ @@ -678,27 +678,27 @@ e.isometry.diam_image s @[simp] lemma preimage_ball (x : E₂) (r : ℝ) : e ⁻¹' (metric.ball x r) = metric.ball (e.symm x) r := -e.to_isometric.preimage_ball x r +e.to_isometry_equiv.preimage_ball x r @[simp] lemma preimage_sphere (x : E₂) (r : ℝ) : e ⁻¹' (metric.sphere x r) = metric.sphere (e.symm x) r := -e.to_isometric.preimage_sphere x r +e.to_isometry_equiv.preimage_sphere x r @[simp] lemma preimage_closed_ball (x : E₂) (r : ℝ) : e ⁻¹' (metric.closed_ball x r) = metric.closed_ball (e.symm x) r := -e.to_isometric.preimage_closed_ball x r +e.to_isometry_equiv.preimage_closed_ball x r @[simp] lemma image_ball (x : E) (r : ℝ) : e '' (metric.ball x r) = metric.ball (e x) r := -e.to_isometric.image_ball x r +e.to_isometry_equiv.image_ball x r @[simp] lemma image_sphere (x : E) (r : ℝ) : e '' (metric.sphere x r) = metric.sphere (e x) r := -e.to_isometric.image_sphere x r +e.to_isometry_equiv.image_sphere x r @[simp] lemma image_closed_ball (x : E) (r : ℝ) : e '' (metric.closed_ball x r) = metric.closed_ball (e x) r := -e.to_isometric.image_closed_ball x r +e.to_isometry_equiv.image_closed_ball x r variables {α : Type*} [topological_space α] diff --git a/src/analysis/normed_space/mazur_ulam.lean b/src/analysis/normed_space/mazur_ulam.lean index 07c94350d1be1..4bae238225025 100644 --- a/src/analysis/normed_space/mazur_ulam.lean +++ b/src/analysis/normed_space/mazur_ulam.lean @@ -13,12 +13,12 @@ import linear_algebra.affine_space.midpoint Mazur-Ulam theorem states that an isometric bijection between two normed affine spaces over `ℝ` is affine. We formalize it in three definitions: -* `isometric.to_real_linear_isometry_equiv_of_map_zero` : given `E ≃ᵢ F` sending `0` to `0`, +* `isometry_equiv.to_real_linear_isometry_equiv_of_map_zero` : given `E ≃ᵢ F` sending `0` to `0`, returns `E ≃ₗᵢ[ℝ] F` with the same `to_fun` and `inv_fun`; -* `isometric.to_real_linear_isometry_equiv` : given `f : E ≃ᵢ F`, returns a linear isometry +* `isometry_equiv.to_real_linear_isometry_equiv` : given `f : E ≃ᵢ F`, returns a linear isometry equivalence `g : E ≃ₗᵢ[ℝ] F` with `g x = f x - f 0`. -* `isometric.to_real_affine_isometry_equiv` : given `f : PE ≃ᵢ PF`, returns an affine isometry - equivalence `g : PE ≃ᵃⁱ[ℝ] PF` whose underlying `isometric` is `f` +* `isometry_equiv.to_real_affine_isometry_equiv` : given `f : PE ≃ᵢ PF`, returns an affine isometry + equivalence `g : PE ≃ᵃⁱ[ℝ] PF` whose underlying `isometry_equiv` is `f` The formalization is based on [Jussi Väisälä, *A Proof of the Mazur-Ulam Theorem*][Vaisala_2003]. @@ -35,7 +35,7 @@ open set affine_map affine_isometry_equiv noncomputable theory -namespace isometric +namespace isometry_equiv include E @@ -48,7 +48,7 @@ begin set z := midpoint ℝ x y, -- Consider the set of `e : E ≃ᵢ E` such that `e x = x` and `e y = y` set s := { e : PE ≃ᵢ PE | e x = x ∧ e y = y }, - haveI : nonempty s := ⟨⟨isometric.refl PE, rfl, rfl⟩⟩, + haveI : nonempty s := ⟨⟨isometry_equiv.refl PE, rfl, rfl⟩⟩, -- On the one hand, `e` cannot send the midpoint `z` of `[x, y]` too far have h_bdd : bdd_above (range $ λ e : s, dist (e z) z), { refine ⟨dist x z + dist x z, forall_range_iff.2 $ subtype.forall.2 _⟩, @@ -59,7 +59,7 @@ begin -- On the other hand, consider the map `f : (E ≃ᵢ E) → (E ≃ᵢ E)` -- sending each `e` to `R ∘ e⁻¹ ∘ R ∘ e`, where `R` is the point reflection in the -- midpoint `z` of `[x, y]`. - set R : PE ≃ᵢ PE := (point_reflection ℝ z).to_isometric, + set R : PE ≃ᵢ PE := (point_reflection ℝ z).to_isometry_equiv, set f : (PE ≃ᵢ PE) → (PE ≃ᵢ PE) := λ e, ((e.trans R).trans e.symm).trans R, -- Note that `f` doubles the value of ``dist (e z) z` have hf_dist : ∀ e, dist (f e z) z = 2 * dist (e z) z, @@ -89,14 +89,14 @@ include F lemma map_midpoint (f : PE ≃ᵢ PF) (x y : PE) : f (midpoint ℝ x y) = midpoint ℝ (f x) (f y) := begin set e : PE ≃ᵢ PE := - ((f.trans $ (point_reflection ℝ $ midpoint ℝ (f x) (f y)).to_isometric).trans f.symm).trans - (point_reflection ℝ $ midpoint ℝ x y).to_isometric, + ((f.trans $ (point_reflection ℝ $ midpoint ℝ (f x) (f y)).to_isometry_equiv).trans f.symm).trans + (point_reflection ℝ $ midpoint ℝ x y).to_isometry_equiv, have hx : e x = x, by simp, have hy : e y = y, by simp, have hm := e.midpoint_fixed hx hy, simp only [e, trans_apply] at hm, - rwa [← eq_symm_apply, to_isometric_symm, point_reflection_symm, coe_to_isometric, - coe_to_isometric, point_reflection_self, symm_apply_eq, point_reflection_fixed_iff] at hm + rwa [← eq_symm_apply, to_isometry_equiv_symm, point_reflection_symm, coe_to_isometry_equiv, + coe_to_isometry_equiv, point_reflection_self, symm_apply_eq, point_reflection_fixed_iff] at hm end /-! @@ -121,7 +121,7 @@ def to_real_linear_isometry_equiv_of_map_zero (f : E ≃ᵢ F) (h0 : f 0 = 0) : /-- **Mazur-Ulam Theorem**: if `f` is an isometric bijection between two normed vector spaces over `ℝ`, then `x ↦ f x - f 0` is a linear isometry equivalence. -/ def to_real_linear_isometry_equiv (f : E ≃ᵢ F) : E ≃ₗᵢ[ℝ] F := -(f.trans (isometric.add_right (f 0)).symm).to_real_linear_isometry_equiv_of_map_zero +(f.trans (isometry_equiv.add_right (f 0)).symm).to_real_linear_isometry_equiv_of_map_zero (by simpa only [sub_eq_add_neg] using sub_self (f 0)) @[simp] lemma to_real_linear_equiv_apply (f : E ≃ᵢ F) (x : E) : @@ -144,7 +144,7 @@ affine_isometry_equiv.mk' f rfl @[simp] lemma coe_to_real_affine_isometry_equiv (f : PE ≃ᵢ PF) : - f.to_real_affine_isometry_equiv.to_isometric = f := + f.to_real_affine_isometry_equiv.to_isometry_equiv = f := by { ext, refl } -end isometric +end isometry_equiv diff --git a/src/measure_theory/function/conditional_expectation/basic.lean b/src/measure_theory/function/conditional_expectation/basic.lean index 4a71620469317..3576925d83dce 100644 --- a/src/measure_theory/function/conditional_expectation/basic.lean +++ b/src/measure_theory/function/conditional_expectation/basic.lean @@ -303,7 +303,7 @@ section complete_subspace /-! ## The subspace `Lp_meas` is complete. -We define an `isometric` between `Lp_meas_subgroup` and the `Lp` space corresponding to the +We define an `isometry_equiv` between `Lp_meas_subgroup` and the `Lp` space corresponding to the measure `μ.trim hm`. As a consequence, the completeness of `Lp` implies completeness of `Lp_meas_subgroup` (and `Lp_meas`). -/ @@ -488,7 +488,7 @@ variables (𝕜) /-- `Lp_meas_subgroup` and `Lp_meas` are isometric. -/ def Lp_meas_subgroup_to_Lp_meas_iso [hp : fact (1 ≤ p)] : Lp_meas_subgroup F m p μ ≃ᵢ Lp_meas F 𝕜 m p μ := -isometric.refl (Lp_meas_subgroup F m p μ) +isometry_equiv.refl (Lp_meas_subgroup F m p μ) /-- `Lp_meas` and `Lp F p (μ.trim hm)` are isometric, with a linear equivalence. -/ def Lp_meas_to_Lp_trim_lie [hp : fact (1 ≤ p)] (hm : m ≤ m0) : diff --git a/src/measure_theory/measure/hausdorff.lean b/src/measure_theory/measure/hausdorff.lean index da68167418b46..0a7554a955e76 100644 --- a/src/measure_theory/measure/hausdorff.lean +++ b/src/measure_theory/measure/hausdorff.lean @@ -387,13 +387,13 @@ lemma isometry_map_mk_metric (m : ℝ≥0∞ → ℝ≥0∞) {f : X → Y} (hf : map f (mk_metric m) = restrict (range f) (mk_metric m) := by rw [← isometry_comap_mk_metric _ hf H, map_comap] -lemma isometric_comap_mk_metric (m : ℝ≥0∞ → ℝ≥0∞) (f : X ≃ᵢ Y) : +lemma isometry_equiv_comap_mk_metric (m : ℝ≥0∞ → ℝ≥0∞) (f : X ≃ᵢ Y) : comap f (mk_metric m) = mk_metric m := isometry_comap_mk_metric _ f.isometry (or.inr f.surjective) -lemma isometric_map_mk_metric (m : ℝ≥0∞ → ℝ≥0∞) (f : X ≃ᵢ Y) : +lemma isometry_equiv_map_mk_metric (m : ℝ≥0∞ → ℝ≥0∞) (f : X ≃ᵢ Y) : map f (mk_metric m) = mk_metric m := -by rw [← isometric_comap_mk_metric _ f, map_comap_of_surjective f.surjective] +by rw [← isometry_equiv_comap_mk_metric _ f, map_comap_of_surjective f.surjective] lemma trim_mk_metric [measurable_space X] [borel_space X] (m : ℝ≥0∞ → ℝ≥0∞) : (mk_metric m : outer_measure X).trim = mk_metric m := @@ -942,7 +942,7 @@ end end isometry -namespace isometric +namespace isometry_equiv @[simp] lemma hausdorff_measure_image (e : X ≃ᵢ Y) (d : ℝ) (s : set X) : μH[d] (e '' s) = μH[d] s := @@ -952,4 +952,4 @@ e.isometry.hausdorff_measure_image (or.inr e.surjective) s μH[d] (e ⁻¹' s) = μH[d] s := by rw [← e.image_symm, e.symm.hausdorff_measure_image] -end isometric +end isometry_equiv diff --git a/src/topology/continuous_function/compact.lean b/src/topology/continuous_function/compact.lean index 159e697800626..191ccc57697a7 100644 --- a/src/topology/continuous_function/compact.lean +++ b/src/topology/continuous_function/compact.lean @@ -83,7 +83,7 @@ When `α` is compact, and `β` is a metric space, the bounded continuous maps ` isometric to `C(α, β)`. -/ @[simps to_equiv apply symm_apply { fully_applied := ff }] -def isometric_bounded_of_compact : +def isometry_equiv_bounded_of_compact : C(α, β) ≃ᵢ (α →ᵇ β) := { isometry_to_fun := λ x y, rfl, to_equiv := equiv_bounded_of_compact α β } @@ -127,11 +127,11 @@ by simp only [← dist_mk_of_compact, dist_lt_iff_of_compact C0, mk_of_compact_a end instance [complete_space β] : complete_space (C(α, β)) := -(isometric_bounded_of_compact α β).complete_space +(isometry_equiv_bounded_of_compact α β).complete_space /-- See also `continuous_map.continuous_eval'` -/ @[continuity] lemma continuous_eval : continuous (λ p : C(α, β) × α, p.1 p.2) := -continuous_eval.comp ((isometric_bounded_of_compact α β).continuous.prod_map continuous_id) +continuous_eval.comp ((isometry_equiv_bounded_of_compact α β).continuous.prod_map continuous_id) /-- See also `continuous_map.continuous_eval_const` -/ @[continuity] lemma continuous_eval_const (x : α) : continuous (λ f : C(α, β), f x) := @@ -259,8 +259,9 @@ rfl @[simp] -lemma linear_isometry_bounded_of_compact_to_isometric : - (linear_isometry_bounded_of_compact α E 𝕜).to_isometric = (isometric_bounded_of_compact α E) := +lemma linear_isometry_bounded_of_compact_to_isometry_equiv : + (linear_isometry_bounded_of_compact α E 𝕜).to_isometry_equiv = + (isometry_equiv_bounded_of_compact α E) := rfl @[simp] diff --git a/src/topology/metric_space/gromov_hausdorff.lean b/src/topology/metric_space/gromov_hausdorff.lean index 092a76e9be846..19158f8e22df5 100644 --- a/src/topology/metric_space/gromov_hausdorff.lean +++ b/src/topology/metric_space/gromov_hausdorff.lean @@ -67,7 +67,7 @@ private def isometry_rel : nonempty_compacts ℓ_infty_ℝ → nonempty_compacts /-- This is indeed an equivalence relation -/ private lemma is_equivalence_isometry_rel : equivalence isometry_rel := -⟨λ x, ⟨isometric.refl _⟩, λ x y ⟨e⟩, ⟨e.symm⟩, λ x y z ⟨e⟩ ⟨f⟩, ⟨e.trans f⟩⟩ +⟨λ x, ⟨isometry_equiv.refl _⟩, λ x y ⟨e⟩, ⟨e.symm⟩, λ x y z ⟨e⟩ ⟨f⟩, ⟨e.trans f⟩⟩ /-- setoid instance identifying two isometric nonempty compact subspaces of ℓ^∞(ℝ) -/ instance isometry_rel.setoid : setoid (nonempty_compacts ℓ_infty_ℝ) := @@ -93,13 +93,13 @@ begin simp only [to_GH_space, quotient.eq], refine ⟨λ h, _, _⟩, { rcases setoid.symm h with ⟨e⟩, - have f := (Kuratowski_embedding.isometry X).isometric_on_range.trans e, + have f := (Kuratowski_embedding.isometry X).isometry_equiv_on_range.trans e, use [λ x, f x, isometry_subtype_coe.comp f.isometry], rw [range_comp, f.range_eq_univ, set.image_univ, subtype.range_coe], refl }, { rintros ⟨Ψ, ⟨isomΨ, rangeΨ⟩⟩, - have f := ((Kuratowski_embedding.isometry X).isometric_on_range.symm.trans - isomΨ.isometric_on_range).symm, + have f := ((Kuratowski_embedding.isometry X).isometry_equiv_on_range.symm.trans + isomΨ.isometry_equiv_on_range).symm, have E : (range Ψ ≃ᵢ nonempty_compacts.Kuratowski_embedding X) = (p ≃ᵢ range (Kuratowski_embedding X)), by { dunfold nonempty_compacts.Kuratowski_embedding, rw [rangeΨ]; refl }, @@ -127,7 +127,7 @@ end /-- Two nonempty compact spaces have the same image in `GH_space` if and only if they are isometric. -/ -lemma to_GH_space_eq_to_GH_space_iff_isometric {X : Type u} [metric_space X] [compact_space X] +lemma to_GH_space_eq_to_GH_space_iff_isometry_equiv {X : Type u} [metric_space X] [compact_space X] [nonempty X] {Y : Type v} [metric_space Y] [compact_space Y] [nonempty Y] : to_GH_space X = to_GH_space Y ↔ nonempty (X ≃ᵢ Y) := ⟨begin @@ -137,15 +137,15 @@ lemma to_GH_space_eq_to_GH_space_iff_isometric {X : Type u} [metric_space X] [co (nonempty_compacts.Kuratowski_embedding Y)) = ((range (Kuratowski_embedding X)) ≃ᵢ (range (Kuratowski_embedding Y))), by { dunfold nonempty_compacts.Kuratowski_embedding, refl }, - have f := (Kuratowski_embedding.isometry X).isometric_on_range, - have g := (Kuratowski_embedding.isometry Y).isometric_on_range.symm, + have f := (Kuratowski_embedding.isometry X).isometry_equiv_on_range, + have g := (Kuratowski_embedding.isometry Y).isometry_equiv_on_range.symm, exact ⟨f.trans $ (cast I e).trans g⟩ end, begin rintro ⟨e⟩, simp only [to_GH_space, quotient.eq], - have f := (Kuratowski_embedding.isometry X).isometric_on_range.symm, - have g := (Kuratowski_embedding.isometry Y).isometric_on_range, + have f := (Kuratowski_embedding.isometry X).isometry_equiv_on_range.symm, + have g := (Kuratowski_embedding.isometry Y).isometry_equiv_on_range, have I : ((range (Kuratowski_embedding X)) ≃ᵢ (range (Kuratowski_embedding Y))) = ((nonempty_compacts.Kuratowski_embedding X) ≃ᵢ (nonempty_compacts.Kuratowski_embedding Y)), @@ -408,9 +408,9 @@ instance : metric_space GH_space := { exact Hausdorff_edist_ne_top_of_nonempty_of_bounded (range_nonempty _) (range_nonempty _) hΦ.bounded hΨ.bounded } }, have T : ((range Ψ) ≃ᵢ y.rep) = ((range Φ) ≃ᵢ y.rep), by rw this, - have eΨ := cast T Ψisom.isometric_on_range.symm, - have e := Φisom.isometric_on_range.trans eΨ, - rw [← x.to_GH_space_rep, ← y.to_GH_space_rep, to_GH_space_eq_to_GH_space_iff_isometric], + have eΨ := cast T Ψisom.isometry_equiv_on_range.symm, + have e := Φisom.isometry_equiv_on_range.trans eΨ, + rw [← x.to_GH_space_rep, ← y.to_GH_space_rep, to_GH_space_eq_to_GH_space_iff_isometry_equiv], exact ⟨e⟩ end, dist_triangle := λ x y z, begin @@ -1016,9 +1016,9 @@ begin have : ∀ n, (X3 n).to_GH_space = u n, { assume n, rw [nonempty_compacts.to_GH_space, ← (u n).to_GH_space_rep, - to_GH_space_eq_to_GH_space_iff_isometric], + to_GH_space_eq_to_GH_space_iff_isometry_equiv], constructor, - convert (isom n).isometric_on_range.symm, }, + convert (isom n).isometry_equiv_on_range.symm, }, -- Finally, we have proved the convergence of `u n` exact ⟨L.to_GH_space, by simpa only [this] using M⟩ end diff --git a/src/topology/metric_space/hausdorff_dimension.lean b/src/topology/metric_space/hausdorff_dimension.lean index 28dd8e7a54eeb..a40f4865e133c 100644 --- a/src/topology/metric_space/hausdorff_dimension.lean +++ b/src/topology/metric_space/hausdorff_dimension.lean @@ -387,7 +387,7 @@ end antilipschitz_with lemma isometry.dimH_image (hf : isometry f) (s : set X) : dimH (f '' s) = dimH s := le_antisymm (hf.lipschitz.dimH_image_le _) (hf.antilipschitz.le_dimH_image _) -namespace isometric +namespace isometry_equiv @[simp] lemma dimH_image (e : X ≃ᵢ Y) (s : set X) : dimH (e '' s) = dimH s := e.isometry.dimH_image s @@ -398,7 +398,7 @@ by rw [← e.image_symm, e.symm.dimH_image] lemma dimH_univ (e : X ≃ᵢ Y) : dimH (univ : set X) = dimH (univ : set Y) := by rw [← e.dimH_preimage univ, preimage_univ] -end isometric +end isometry_equiv namespace continuous_linear_equiv diff --git a/src/topology/metric_space/isometry.lean b/src/topology/metric_space/isometry.lean index 37e013925d5fe..32b3f5e4bd66d 100644 --- a/src/topology/metric_space/isometry.lean +++ b/src/topology/metric_space/isometry.lean @@ -233,13 +233,13 @@ end /-- `α` and `β` are isometric if there is an isometric bijection between them. -/ @[nolint has_nonempty_instance] -- such a bijection need not exist -structure isometric (α : Type*) (β : Type*) [pseudo_emetric_space α] [pseudo_emetric_space β] +structure isometry_equiv (α β : Type*) [pseudo_emetric_space α] [pseudo_emetric_space β] extends α ≃ β := (isometry_to_fun : isometry to_fun) -infix ` ≃ᵢ `:25 := isometric +infix ` ≃ᵢ `:25 := isometry_equiv -namespace isometric +namespace isometry_equiv section pseudo_emetric_space variables [pseudo_emetric_space α] [pseudo_emetric_space β] [pseudo_emetric_space γ] @@ -310,7 +310,7 @@ def simps.apply (h : α ≃ᵢ β) : α → β := h /-- See Note [custom simps projection] -/ def simps.symm_apply (h : α ≃ᵢ β) : β → α := h.symm -initialize_simps_projections isometric +initialize_simps_projections isometry_equiv (to_equiv_to_fun → apply, to_equiv_inv_fun → symm_apply) @[simp] lemma symm_symm (h : α ≃ᵢ β) : h.symm.symm = h := to_equiv_inj h.to_equiv.symm_symm @@ -394,9 +394,9 @@ h.to_homeomorph.comp_continuous_iff' /-- The group of isometries. -/ instance : group (α ≃ᵢ α) := - { one := isometric.refl _, + { one := isometry_equiv.refl _, mul := λ e₁ e₂, e₂.trans e₁, - inv := isometric.symm, + inv := isometry_equiv.symm, mul_assoc := λ e₁ e₂ e₃, rfl, one_mul := λ e, ext $ λ _, rfl, mul_one := λ e, ext $ λ _, rfl, @@ -414,7 +414,7 @@ lemma mul_apply (e₁ e₂ : α ≃ᵢ α) (x : α) : (e₁ * e₂) x = e₁ (e protected lemma complete_space [complete_space β] (e : α ≃ᵢ β) : complete_space α := complete_space_of_is_complete_univ $ is_complete_of_complete_image e.isometry.uniform_inducing $ - by rwa [set.image_univ, isometric.range_eq_univ, ← complete_space_iff_is_complete_univ] + by rwa [set.image_univ, isometry_equiv.range_eq_univ, ← complete_space_iff_is_complete_univ] lemma complete_space_iff (e : α ≃ᵢ β) : complete_space α ↔ complete_space β := by { split; introI H, exacts [e.symm.complete_space, e.complete_space] } @@ -460,12 +460,12 @@ by rw [← h.preimage_symm, h.symm.preimage_closed_ball, symm_symm] end pseudo_metric_space -end isometric +end isometry_equiv /-- An isometry induces an isometric isomorphism between the source space and the range of the isometry. -/ @[simps to_equiv apply { simp_rhs := tt }] -def isometry.isometric_on_range [emetric_space α] [pseudo_emetric_space β] {f : α → β} +def isometry.isometry_equiv_on_range [emetric_space α] [pseudo_emetric_space β] {f : α → β} (h : isometry f) : α ≃ᵢ range f := { isometry_to_fun := λx y, by simpa [subtype.edist_eq] using h x y, to_equiv := equiv.of_injective f h.injective } From c0439b4877c24a117bfdd9e32faf62eee9b115eb Mon Sep 17 00:00:00 2001 From: Niels Voss Date: Tue, 17 Jan 2023 01:41:19 +0000 Subject: [PATCH 0265/1291] feat(number_theory): prove the existence of infinitely many Fermat pseudoprimes to any base (#17632) A natural number `n` is a probable prime to base `b` if `n` divides `b ^ (n - 1) - 1`. `n` is a Fermat pseudoprime to base `b` if it is a composite probable prime. This commit defines Fermat pseudoprimes and proves that for any positive base `b`, there exist an infinite number of Fermat pseudoprimes to that base. Co-authored-by: Niels Voss --- src/number_theory/fermat_psp.lean | 421 ++++++++++++++++++++++++++++++ 1 file changed, 421 insertions(+) create mode 100644 src/number_theory/fermat_psp.lean diff --git a/src/number_theory/fermat_psp.lean b/src/number_theory/fermat_psp.lean new file mode 100644 index 0000000000000..a1860ef1866ea --- /dev/null +++ b/src/number_theory/fermat_psp.lean @@ -0,0 +1,421 @@ +/- +Copyright (c) 2022 Niels Voss. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Niels Voss +-/ +import data.nat.prime +import field_theory.finite.basic +import order.filter.cofinite + +/-! +# Fermat Pseudoprimes + +In this file we define Fermat pseudoprimes: composite numbers that pass the Fermat primality test. +A natural number `n` passes the Fermat primality test to base `b` (and is therefore deemed a +"probable prime") if `n` divides `b ^ (n - 1) - 1`. `n` is a Fermat pseudoprime to base `b` if `n` +is a composite number that passes the Fermat primality test to base `b` and is coprime with `b`. + +Fermat pseudoprimes can also be seen as composite numbers for which Fermat's little theorem holds +true. + +Numbers which are Fermat pseudoprimes to all bases are known as Carmichael numbers (not yet defined +in this file). + +## Main Results + +The main definitions for this file are + +- `fermat_psp.probable_prime`: A number `n` is a probable prime to base `b` if it passes the Fermat + primality test; that is, if `n` divides `b ^ (n - 1) - 1` +- `fermat_psp`: A number `n` is a pseudoprime to base `b` if it is a probable prime to base `b`, is + composite, and is coprime with `b` (this last condition is automatically true if `n` divides + `b ^ (n - 1) - 1`, but some sources include it in the definition). + +Note that all composite numbers are pseudoprimes to base 0 and 1, and that the definiton of +`probable_prime` in this file implies that all numbers are probable primes to bases 0 and 1, and +that 0 and 1 are probable primes to any base. + +The main theorems are +- `fermat_psp.exists_infinite_pseudoprimes`: there are infinite pseudoprimes to any base `b ≥ 1` +-/ + +/-- +`n` is a probable prime to base `b` if `n` passes the Fermat primality test; that is, `n` divides +`b ^ (n - 1) - 1`. +This definition implies that all numbers are probable primes to base 0 or 1, and that 0 and 1 are +probable primes to any base. +-/ +def fermat_psp.probable_prime (n b : ℕ) : Prop := n ∣ b ^ (n - 1) - 1 + +/-- +`n` is a Fermat pseudoprime to base `b` if `n` is a probable prime to base `b` and is composite. By +this definition, all composite natural numbers are pseudoprimes to base 0 and 1. This definition +also permits `n` to be less than `b`, so that 4 is a pseudoprime to base 5, for example. +-/ +def fermat_psp (n b : ℕ) : Prop := fermat_psp.probable_prime n b ∧ ¬n.prime ∧ 1 < n + +namespace fermat_psp + +instance decidable_probable_prime (n b : ℕ) : decidable (probable_prime n b) := +nat.decidable_dvd _ _ + +instance decidable_psp (n b : ℕ) : decidable (fermat_psp n b) := and.decidable + +/-- +If `n` passes the Fermat primality test to base `b`, then `n` is coprime with `b`, assuming that +`n` and `b` are both positive. +-/ +lemma coprime_of_probable_prime {n b : ℕ} (h : probable_prime n b) (h₁ : 1 ≤ n) (h₂ : 1 ≤ b) : + nat.coprime n b := +begin + by_cases h₃ : 2 ≤ n, + + { -- To prove that `n` is coprime with `b`, we we need to show that for all prime factors of `n`, + -- we can derive a contradiction if `n` divides `b`. + apply nat.coprime_of_dvd, + + -- If `k` is a prime number that divides both `n` and `b`, then we know that `n = m * k` and + -- `b = j * k` for some natural numbers `m` and `j`. We substitute these into the hypothesis. + rintros k hk ⟨m, rfl⟩ ⟨j, rfl⟩, + + -- Because prime numbers do not divide 1, it suffices to show that `k ∣ 1` to prove a + -- contradiction + apply nat.prime.not_dvd_one hk, + + -- Since `n` divides `b ^ (n - 1) - 1`, `k` also divides `b ^ (n - 1) - 1` + replace h := dvd_of_mul_right_dvd h, + + -- Because `k` divides `b ^ (n - 1) - 1`, if we can show that `k` also divides `b ^ (n - 1)`, + -- then we know `k` divides 1. + rw [nat.dvd_add_iff_right h, nat.sub_add_cancel (nat.one_le_pow _ _ h₂)], + + -- Since `k` divides `b`, `k` also divides any power of `b` except `b ^ 0`. Therefore, it + -- suffices to show that `n - 1` isn't zero. However, we know that `n - 1` isn't zero because we + -- assumed `2 ≤ n` when doing `by_cases`. + refine dvd_of_mul_right_dvd (dvd_pow_self (k * j) _), + linarith }, + + -- If `n = 1`, then it follows trivially that `n` is coprime with `b`. + { rw (show n = 1, by linarith), + norm_num } +end + +lemma probable_prime_iff_modeq (n : ℕ) {b : ℕ} (h : 1 ≤ b) : + probable_prime n b ↔ b ^ (n - 1) ≡ 1 [MOD n] := +begin + have : 1 ≤ b ^ (n - 1) := one_le_pow_of_one_le h (n - 1), -- For exact_mod_cast + rw nat.modeq.comm, + split, + { intro h₁, + apply nat.modeq_of_dvd, + exact_mod_cast h₁, }, + { intro h₁, + exact_mod_cast nat.modeq.dvd h₁, }, +end + +/-- +If `n` is a Fermat pseudoprime to base `b`, then `n` is coprime with `b`, assuming that `b` is +positive. + +This lemma is a small wrapper based on `coprime_of_probable_prime` +-/ +lemma coprime_of_fermat_psp {n b : ℕ} (h : fermat_psp n b) (h₁ : 1 ≤ b) : nat.coprime n b := +begin + rcases h with ⟨hp, hn₁, hn₂⟩, + exact coprime_of_probable_prime hp (by linarith) h₁, +end + +/-- +All composite numbers are Fermat pseudoprimes to base 1. +-/ +lemma base_one {n : ℕ} (h₁ : 1 < n) (h₂ : ¬n.prime) : fermat_psp n 1 := +begin + refine ⟨show n ∣ 1 ^ (n - 1) - 1, from _, h₂, h₁⟩, + exact (show 0 = 1 ^ (n - 1) - 1, by norm_num) ▸ dvd_zero n, +end + +-- Lemmas that are needed to prove statements in this file, but aren't directly related to Fermat +-- pseudoprimes +section helper_lemmas + +private lemma pow_gt_exponent {a : ℕ} (b : ℕ) (h : 2 ≤ a) : b < a ^ b := +lt_of_lt_of_le (nat.lt_two_pow b) $ nat.pow_le_pow_of_le_left h _ + +private lemma a_id_helper {a b : ℕ} (ha : 2 ≤ a) (hb : 2 ≤ b) : 2 ≤ (a ^ b - 1) / (a - 1) := +begin + change 1 < _, + have h₁ : a - 1 ∣ a ^ b - 1 := by simpa only [one_pow] using nat_sub_dvd_pow_sub_pow a 1 b, + rw [nat.lt_div_iff_mul_lt h₁, mul_one, tsub_lt_tsub_iff_right (nat.le_of_succ_le ha)], + convert pow_lt_pow (nat.lt_of_succ_le ha) hb, + rw pow_one +end + +private lemma b_id_helper {a b : ℕ} (ha : 2 ≤ a) (hb : 2 < b) : 2 ≤ (a ^ b + 1) / (a + 1) := +begin + rw nat.le_div_iff_mul_le (nat.zero_lt_succ _), + apply nat.succ_le_succ, + calc 2 * a + 1 ≤ a ^ 2 * a : by nlinarith + ... = a ^ 3 : by rw pow_succ' a 2 + ... ≤ a ^ b : pow_le_pow (nat.le_of_succ_le ha) hb +end + +private lemma AB_id_helper (b p : ℕ) (hb : 2 ≤ b) (hp : odd p) + : (b ^ p - 1) / (b - 1) * ((b ^ p + 1) / (b + 1)) = (b ^ (2 * p) - 1) / (b ^ 2 - 1) := +begin + have q₁ : b - 1 ∣ b ^ p - 1 := by simpa only [one_pow] using nat_sub_dvd_pow_sub_pow b 1 p, + have q₂ : b + 1 ∣ b ^ p + 1 := by simpa only [one_pow] using hp.nat_add_dvd_pow_add_pow b 1, + convert nat.div_mul_div_comm q₁ q₂; rw [mul_comm (_ - 1), ←nat.sq_sub_sq], + { ring_exp }, + { simp } +end + +/-- +Used in the proof of `psp_from_prime_psp` +-/ +private lemma bp_helper {b p : ℕ} (hb : 0 < b) (hp : 1 ≤ p) : + b ^ (2 * p) - 1 - (b ^ 2 - 1) = b * (b ^ (p - 1) - 1) * (b ^ p + b) := +have hi_bsquared : 1 ≤ b ^ 2 := nat.one_le_pow _ _ hb, +calc b ^ (2 * p) - 1 - (b ^ 2 - 1) = b ^ (2 * p) - (1 + (b ^ 2 - 1)) : by rw nat.sub_sub + ... = b ^ (2 * p) - (1 + b ^ 2 - 1) : by rw nat.add_sub_assoc hi_bsquared + ... = b ^ (2 * p) - (b ^ 2) : by rw nat.add_sub_cancel_left + ... = b ^ (p * 2) - (b ^ 2) : by rw mul_comm + ... = (b ^ p) ^ 2 - (b ^ 2) : by rw pow_mul + ... = (b ^ p + b) * (b ^ p - b) : by rw nat.sq_sub_sq + ... = (b ^ p - b) * (b ^ p + b) : by rw mul_comm + ... = (b ^ (p - 1 + 1) - b) * (b ^ p + b) : by rw nat.sub_add_cancel hp + ... = (b * b ^ (p - 1) - b) * (b ^ p + b) : by rw pow_succ + ... = (b * b ^ (p - 1) - b * 1) * (b ^ p + b) : by rw mul_one + ... = b * (b ^ (p - 1) - 1) * (b ^ p + b) : by rw nat.mul_sub_left_distrib + +end helper_lemmas + +/-- +Given a prime `p` which does not divide `b * (b ^ 2 - 1)`, we can produce a number `n` which is +larger than `p` and pseudoprime to base `b`. We do this by defining +`n = ((b ^ p - 1) / (b - 1)) * ((b ^ p + 1) / (b + 1))` + +The primary purpose of this definition is to help prove `exists_infinite_pseudoprimes`. For a proof +that `n` is actually pseudoprime to base `b`, see `psp_from_prime_psp`, and for a proof that `n` is +greater than `p`, see `psp_from_prime_gt_p`. + +This lemma is intended to be used when `2 ≤ b`, `2 < p`, `p` is prime, and `¬p ∣ b * (b ^ 2 - 1)`, +because those are the hypotheses for `psp_from_prime_psp`. +-/ +private def psp_from_prime (b : ℕ) (p : ℕ) : ℕ := (b ^ p - 1) / (b - 1) * ((b ^ p + 1) / (b + 1)) + +/-- +This is a proof that the number produced using `psp_from_prime` is actually pseudoprime to base `b`. +The primary purpose of this lemma is to help prove `exists_infinite_pseudoprimes`. + +We use as a rough outline of the proof. +-/ +private lemma psp_from_prime_psp {b : ℕ} (b_ge_two : 2 ≤ b) {p : ℕ} (p_prime : p.prime) + (p_gt_two : 2 < p) (not_dvd : ¬p ∣ b * (b ^ 2 - 1)) : + fermat_psp (psp_from_prime b p) b := +begin + unfold psp_from_prime, + set A := (b ^ p - 1) / (b - 1), + set B := (b ^ p + 1) / (b + 1), + + -- Inequalities + have hi_A : 1 < A := a_id_helper (nat.succ_le_iff.mp b_ge_two) (nat.prime.one_lt p_prime), + have hi_B : 1 < B := b_id_helper (nat.succ_le_iff.mp b_ge_two) p_gt_two, + have hi_AB : 1 < A * B := one_lt_mul'' hi_A hi_B, + have hi_b : 0 < b := by linarith, + have hi_p : 1 ≤ p := nat.one_le_of_lt p_gt_two, + have hi_bsquared : 0 < b ^ 2 - 1 := by nlinarith [nat.one_le_pow 2 b hi_b], + have hi_bpowtwop : 1 ≤ b ^ (2 * p) := nat.one_le_pow (2 * p) b hi_b, + have hi_bpowpsubone : 1 ≤ b ^ (p - 1) := nat.one_le_pow (p - 1) b hi_b, + + -- Other useful facts + have p_odd : odd p := p_prime.odd_of_ne_two p_gt_two.ne.symm, + have AB_not_prime : ¬nat.prime (A * B) := nat.not_prime_mul hi_A hi_B, + have AB_id : A * B = (b ^ (2 * p) - 1) / (b ^ 2 - 1) := AB_id_helper _ _ b_ge_two p_odd, + have hd : b ^ 2 - 1 ∣ b ^ (2 * p) - 1, + { simpa only [one_pow, pow_mul] using nat_sub_dvd_pow_sub_pow _ 1 p }, + + -- We know that `A * B` is not prime, and that `1 < A * B`. Since two conditions of being + -- pseudoprime are satisfied, we only need to show that `A * B` is probable prime to base `b` + refine ⟨_, AB_not_prime, hi_AB⟩, + + -- Used to prove that `2 * p * (b ^ 2 - 1) ∣ (b ^ 2 - 1) * (A * B - 1)`. + have ha₁ : (b ^ 2 - 1) * (A * B - 1) = b * (b ^ (p - 1) - 1) * (b ^ p + b), + { apply_fun (λ x, x * (b ^ 2 - 1)) at AB_id, + rw nat.div_mul_cancel hd at AB_id, + apply_fun (λ x, x - (b ^ 2 - 1)) at AB_id, + nth_rewrite 1 ←one_mul (b ^ 2 - 1) at AB_id, + rw [←nat.mul_sub_right_distrib, mul_comm] at AB_id, + rw AB_id, + exact bp_helper hi_b hi_p }, + -- If `b` is even, then `b^p` is also even, so `2 ∣ b^p + b` + -- If `b` is odd, then `b^p` is also odd, so `2 ∣ b^p + b` + have ha₂ : 2 ∣ b ^ p + b, + { by_cases h : even b, + { replace h : 2 ∣ b := even_iff_two_dvd.mp h, + have : p ≠ 0 := by linarith, + have : 2 ∣ b^p := dvd_pow h this, + exact dvd_add this h }, + { have h : odd b := nat.odd_iff_not_even.mpr h, + have : odd (b ^ p) := odd.pow h, + have : even (b ^ p + b) := odd.add_odd this h, + exact even_iff_two_dvd.mp this } }, + -- Since `b` isn't divisible by `p`, `b` is coprime with `p`. we can use Fermat's Little Theorem + -- to prove this. + have ha₃ : p ∣ b ^ (p - 1) - 1, + { have : ¬p ∣ b := mt (assume h : p ∣ b, dvd_mul_of_dvd_left h _) not_dvd, + have : p.coprime b := or.resolve_right (nat.coprime_or_dvd_of_prime p_prime b) this, + have : is_coprime (b : ℤ) ↑p := this.symm.is_coprime, + have : ↑b ^ (p - 1) ≡ 1 [ZMOD ↑p] := int.modeq.pow_card_sub_one_eq_one p_prime this, + have : ↑p ∣ ↑b ^ (p - 1) - ↑1 := int.modeq.dvd (int.modeq.symm this), + exact_mod_cast this }, + -- Because `p - 1` is even, there is a `c` such that `2 * c = p - 1`. `nat_sub_dvd_pow_sub_pow` + -- implies that `b ^ c - 1 ∣ (b ^ c) ^ 2 - 1`, and `(b ^ c) ^ 2 = b ^ (p - 1)`. + have ha₄ : b ^ 2 - 1 ∣ b ^ (p - 1) - 1, + { cases p_odd with k hk, + have : 2 ∣ p - 1 := ⟨k, by simp [hk]⟩, + cases this with c hc, + have : b ^ 2 - 1 ∣ (b ^ 2) ^ c - 1 := + by simpa only [one_pow] using nat_sub_dvd_pow_sub_pow _ 1 c, + have : b ^ 2 - 1 ∣ b ^ (2 * c) - 1 := by rwa ←pow_mul at this, + rwa ←hc at this }, + -- Used to prove that `2 * p` divides `A * B - 1` + have ha₅ : 2 * p * (b ^ 2 - 1) ∣ (b ^ 2 - 1) * (A * B - 1), + { suffices q : 2 * p * (b ^ 2 - 1) ∣ b * (b ^ (p - 1) - 1) * (b ^ p + b), + { rwa ha₁ }, + -- We already proved that `b ^ 2 - 1 ∣ b ^ (p - 1) - 1`. + -- Since `2 ∣ b ^ p + b` and `p ∣ b ^ p + b`, if we show that 2 and p are coprime, then we + -- know that `2 * p ∣ b ^ p + b` + have q₁ : nat.coprime p (b ^ 2 - 1), + { have q₂ : ¬p ∣ b ^ 2 - 1, + { rw mul_comm at not_dvd, + exact mt (assume h : p ∣ b ^ 2 - 1, dvd_mul_of_dvd_left h _) not_dvd }, + exact (nat.prime.coprime_iff_not_dvd p_prime).mpr q₂ }, + have q₂ : p * (b ^ 2 - 1) ∣ b ^ (p - 1) - 1 := nat.coprime.mul_dvd_of_dvd_of_dvd q₁ ha₃ ha₄, + have q₃ : p * (b ^ 2 - 1) * 2 ∣ (b ^ (p - 1) - 1) * (b ^ p + b) := mul_dvd_mul q₂ ha₂, + have q₄ : p * (b ^ 2 - 1) * 2 ∣ b * ((b ^ (p - 1) - 1) * (b ^ p + b)), + from dvd_mul_of_dvd_right q₃ _, + rwa [mul_assoc, mul_comm, mul_assoc b] }, + have ha₆ : 2 * p ∣ A * B - 1, + { rw mul_comm at ha₅, + exact nat.dvd_of_mul_dvd_mul_left hi_bsquared ha₅ }, + -- `A * B` divides `b ^ (2 * p) - 1` because `A * B * (b ^ 2 - 1) = b ^ (2 * p) - 1`. + -- This can be proven by multiplying both sides of `AB_id` by `b ^ 2 - 1`. + have ha₇ : A * B ∣ b ^ (2 * p) - 1, + { use b ^ 2 - 1, + have : A * B * (b ^ 2 - 1) = (b ^ (2 * p) - 1) / (b ^ 2 - 1) * (b ^ 2 - 1), + from congr_arg (λ x : ℕ, x * (b ^ 2 - 1)) AB_id, + simpa only [add_comm, nat.div_mul_cancel hd, nat.sub_add_cancel hi_bpowtwop] using this.symm }, + -- Since `2 * p ∣ A * B - 1`, there is a number `q` such that `2 * p * q = A * B - 1`. + -- By `nat_sub_dvd_pow_sub_pow`, we know that `b ^ (2 * p) - 1 ∣ b ^ (2 * p * q) - 1`. + -- This means that `b ^ (2 * p) - 1 ∣ b ^ (A * B - 1) - 1`. + cases ha₆ with q hq, + have ha₈ : b ^ (2 * p) - 1 ∣ b ^ (A * B - 1) - 1 := + by simpa only [one_pow, pow_mul, hq] using nat_sub_dvd_pow_sub_pow _ 1 q, + -- We have proved that `A * B ∣ b ^ (2 * p) - 1` and `b ^ (2 * p) - 1 ∣ b ^ (A * B - 1) - 1`. + -- Therefore, `A * B ∣ b ^ (A * B - 1) - 1`. + exact dvd_trans ha₇ ha₈ +end + +/-- +This is a proof that the number produced using `psp_from_prime` is greater than the prime `p` used +to create it. The primary purpose of this lemma is to help prove `exists_infinite_pseudoprimes`. +-/ +private lemma psp_from_prime_gt_p {b : ℕ} (b_ge_two : 2 ≤ b) {p : ℕ} (p_prime : p.prime) + (p_gt_two : 2 < p) : + p < psp_from_prime b p := +begin + unfold psp_from_prime, + set A := (b ^ p - 1) / (b - 1), + set B := (b ^ p + 1) / (b + 1), + rw show A * B = (b ^ (2 * p) - 1) / (b ^ 2 - 1), + from AB_id_helper _ _ b_ge_two (p_prime.odd_of_ne_two p_gt_two.ne.symm), + have AB_dvd : b ^ 2 - 1 ∣ b ^ (2 * p) - 1, + by simpa only [one_pow, pow_mul] using nat_sub_dvd_pow_sub_pow _ 1 p, + + suffices h : p * (b ^ 2 - 1) < b ^ (2 * p) - 1, + { have h₁ : (p * (b ^ 2 - 1)) / (b ^ 2 - 1) < (b ^ (2 * p) - 1) / (b ^ 2 - 1), + from nat.div_lt_div_of_lt_of_dvd AB_dvd h, + have h₂ : 0 < b ^ 2 - 1, + by linarith [show 3 ≤ b ^ 2 - 1, from le_tsub_of_add_le_left (show 4 ≤ b ^ 2, by nlinarith)], + rwa nat.mul_div_cancel _ h₂ at h₁ }, + + rw [nat.mul_sub_left_distrib, mul_one, pow_mul], + nth_rewrite_rhs 0 ←nat.sub_add_cancel (show 1 ≤ p, by linarith), + rw pow_succ (b ^ 2), + suffices h : p * b ^ 2 < b ^ 2 * (b ^ 2) ^ (p - 1), + { apply gt_of_ge_of_gt, + { exact tsub_le_tsub_left (show 1 ≤ p, by linarith) (b ^ 2 * (b ^ 2) ^ (p - 1)) }, + { have : p ≤ p * b ^ 2 := nat.le_mul_of_pos_right (show 0 < b ^ 2, by nlinarith), + exact tsub_lt_tsub_right_of_le this h } }, + + suffices h : p < (b ^ 2) ^ (p - 1), + { rw mul_comm (b ^ 2), + have : 4 ≤ b ^ 2 := by nlinarith, + have : 0 < b ^ 2 := by linarith, + exact mul_lt_mul_of_pos_right h this }, + + rw [←pow_mul, nat.mul_sub_left_distrib, mul_one], + have : 2 ≤ 2 * p - 2 := le_tsub_of_add_le_left (show 4 ≤ 2 * p, by linarith), + have : 2 + p ≤ 2 * p := by linarith, + have : p ≤ 2 * p - 2 := le_tsub_of_add_le_left this, + exact nat.lt_of_le_of_lt this (pow_gt_exponent _ b_ge_two) +end + +/-- +For all positive bases, there exist Fermat infinite pseudoprimes to that base. +Given in this form: for all numbers `b ≥ 1` and `m`, there exists a pseudoprime `n` to base `b` such +that `m ≤ n`. This form is similar to `nat.exists_infinite_primes`. +-/ +theorem exists_infinite_pseudoprimes {b : ℕ} (h : 1 ≤ b) (m : ℕ) : ∃ n : ℕ, fermat_psp n b ∧ m ≤ n +:= +begin + by_cases b_ge_two : 2 ≤ b, + -- If `2 ≤ b`, then because there exist infinite prime numbers, there is a prime number p such + -- `m ≤ p` and `¬p ∣ b*(b^2 - 1)`. We pick a prime number `b*(b^2 - 1) + 1 + m ≤ p` because we + -- automatically know that `p` is greater than m and that it does not divide `b*(b^2 - 1)` + -- (because `p` can't divide a number less than `p`). + -- From `p`, we can use the lemmas we proved earlier to show that + -- `((b^p - 1)/(b - 1)) * ((b^p + 1)/(b + 1))` is a pseudoprime to base `b`. + { have h := nat.exists_infinite_primes (b * (b ^ 2 - 1) + 1 + m), + cases h with p hp, + cases hp with hp₁ hp₂, + have h₁ : 0 < b := pos_of_gt (nat.succ_le_iff.mp b_ge_two), + have h₂ : 4 ≤ b ^ 2 := pow_le_pow_of_le_left' b_ge_two 2, + have h₃ : 0 < b ^ 2 - 1 := tsub_pos_of_lt (gt_of_ge_of_gt h₂ (by norm_num)), + have h₄ : 0 < b * (b ^ 2 - 1) := mul_pos h₁ h₃, + have h₅ : b * (b ^ 2 - 1) < p := by linarith, + have h₆ : ¬p ∣ b * (b ^ 2 - 1) := nat.not_dvd_of_pos_of_lt h₄ h₅, + have h₇ : b ≤ b * (b ^ 2 - 1) := nat.le_mul_of_pos_right h₃, + have h₈ : 2 ≤ b * (b ^ 2 - 1) := le_trans b_ge_two h₇, + have h₉ : 2 < p := gt_of_gt_of_ge h₅ h₈, + have h₁₀ := psp_from_prime_gt_p b_ge_two hp₂ h₉, + use psp_from_prime b p, + split, + { exact psp_from_prime_psp b_ge_two hp₂ h₉ h₆ }, + { exact le_trans (show m ≤ p, by linarith) (le_of_lt h₁₀) } }, + -- If `¬2 ≤ b`, then `b = 1`. Since all composite numbers are pseudoprimes to base 1, we can pick + -- any composite number greater than m. We choose `2 * (m + 2)` because it is greater than `m` and + -- is composite for all natural numbers `m`. + { have h₁ : b = 1 := by linarith, + rw h₁, + use 2 * (m + 2), + have : ¬nat.prime (2 * (m + 2)) := nat.not_prime_mul (by norm_num) (by norm_num), + exact ⟨base_one (by linarith) this, by linarith⟩ } +end + +theorem frequently_at_top_fermat_psp {b : ℕ} (h : 1 ≤ b) : ∃ᶠ n in filter.at_top, fermat_psp n b := +begin + -- Based on the proof of `nat.frequently_at_top_modeq_one` + refine filter.frequently_at_top.2 (λ n, _), + obtain ⟨p, hp⟩ := exists_infinite_pseudoprimes h n, + exact ⟨p, hp.2, hp.1⟩ +end + +/-- +Infinite set variant of `exists_infinite_pseudoprimes` +-/ +theorem infinite_set_of_prime_modeq_one {b : ℕ} (h : 1 ≤ b) : + set.infinite {n : ℕ | fermat_psp n b} := +nat.frequently_at_top_iff_infinite.mp (frequently_at_top_fermat_psp h) + +end fermat_psp From 48024901a8e2a462363650c50d62248a77cbcab3 Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Tue, 17 Jan 2023 06:52:20 +0000 Subject: [PATCH 0266/1291] chore(analysis/convex/exposed): downgrade imports (#18186) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Everything in this file was stated for ```lean [normed_linear_ordered_field 𝕜] ``` but would have worked for ```lean [topological_space 𝕜] [linear_ordered_ring 𝕜] [order_topology 𝕜] ``` and I have generalized even further where it was easy. --- src/analysis/convex/exposed.lean | 125 ++++++++++++++++++++----------- 1 file changed, 81 insertions(+), 44 deletions(-) diff --git a/src/analysis/convex/exposed.lean b/src/analysis/convex/exposed.lean index ce6b1a7065137..65bd344a1ae25 100644 --- a/src/analysis/convex/exposed.lean +++ b/src/analysis/convex/exposed.lean @@ -5,7 +5,8 @@ Authors: Yaël Dillies, Bhavik Mehta -/ import analysis.convex.extreme import analysis.convex.function -import analysis.normed.order.basic +import topology.algebra.module.basic +import topology.order.basic /-! # Exposed sets @@ -45,15 +46,23 @@ More not-yet-PRed stuff is available on the branch `sperner_again`. open_locale classical affine big_operators open set -variables (𝕜 : Type*) {E : Type*} [normed_linear_ordered_field 𝕜] [add_comm_monoid E] [module 𝕜 E] - [topological_space E] {l : E →L[𝕜] 𝕜} {A B C : set E} {X : finset E} {x : E} +section preorder_semiring + +variables (𝕜 : Type*) {E : Type*} [topological_space 𝕜] [semiring 𝕜] [preorder 𝕜] + [add_comm_monoid E] [topological_space E] [module 𝕜 E] {A B : set E} /-- A set `B` is exposed with respect to `A` iff it maximizes some functional over `A` (and contains all points maximizing it). Written `is_exposed 𝕜 A B`. -/ def is_exposed (A B : set E) : Prop := B.nonempty → ∃ l : E →L[𝕜] 𝕜, B = {x ∈ A | ∀ y ∈ A, l y ≤ l x} -variables {𝕜} +end preorder_semiring + +section ordered_ring + +variables {𝕜 : Type*} {E : Type*} [topological_space 𝕜] [ordered_ring 𝕜] + [add_comm_monoid E] [topological_space E] [module 𝕜 E] + {l : E →L[𝕜] 𝕜} {A B C : set E} {X : finset E} {x : E} /-- A useful way to build exposed sets from intersecting `A` with halfspaces (modelled by an inequality with a functional). -/ @@ -95,25 +104,36 @@ begin (λ x hx, ⟨hBA hx.1, λ y hy, (hw.2 y hy).trans (hx.2 w (hCB hw))⟩)⟩, end -/-- If `B` is an exposed subset of `A`, then `B` is the intersection of `A` with some closed +/-- If `B` is a nonempty exposed subset of `A`, then `B` is the intersection of `A` with some closed halfspace. The converse is *not* true. It would require that the corresponding open halfspace doesn't intersect `A`. -/ -lemma eq_inter_halfspace (hAB : is_exposed 𝕜 A B) : +lemma eq_inter_halfspace' {A B : set E} (hAB : is_exposed 𝕜 A B) (hB : B.nonempty) : ∃ l : E →L[𝕜] 𝕜, ∃ a, B = {x ∈ A | a ≤ l x} := begin - obtain hB | hB := B.eq_empty_or_nonempty, - { refine ⟨0, 1, _⟩, - rw [hB, eq_comm, eq_empty_iff_forall_not_mem], - rintro x ⟨-, h⟩, - rw continuous_linear_map.zero_apply at h, - linarith }, obtain ⟨l, rfl⟩ := hAB hB, obtain ⟨w, hw⟩ := hB, exact ⟨l, l w, subset.antisymm (λ x hx, ⟨hx.1, hx.2 w hw.1⟩) (λ x hx, ⟨hx.1, λ y hy, (hw.2 y hy).trans hx.2⟩)⟩, end -protected lemma inter (hB : is_exposed 𝕜 A B) (hC : is_exposed 𝕜 A C) : +/-- For nontrivial `𝕜`, if `B` is an exposed subset of `A`, then `B` is the intersection of `A` with +some closed halfspace. The converse is *not* true. It would require that the corresponding open +halfspace doesn't intersect `A`. -/ +lemma eq_inter_halfspace [nontrivial 𝕜] {A B : set E} (hAB : is_exposed 𝕜 A B) : + ∃ l : E →L[𝕜] 𝕜, ∃ a, B = {x ∈ A | a ≤ l x} := +begin + obtain rfl | hB := B.eq_empty_or_nonempty, + { refine ⟨0, 1, _⟩, + rw [eq_comm, eq_empty_iff_forall_not_mem], + rintro x ⟨-, h⟩, + rw continuous_linear_map.zero_apply at h, + have : ¬ ((1:𝕜) ≤ 0) := not_le_of_lt zero_lt_one, + contradiction }, + exact hAB.eq_inter_halfspace' hB, +end + +protected lemma inter [has_continuous_add 𝕜] {A B C : set E} (hB : is_exposed 𝕜 A B) + (hC : is_exposed 𝕜 A C) : is_exposed 𝕜 A (B ∩ C) := begin rintro ⟨w, hwB, hwC⟩, @@ -130,7 +150,7 @@ begin (hx w hwB.1)) } end -lemma sInter {F : finset (set E)} (hF : F.nonempty) +lemma sInter [has_continuous_add 𝕜] {F : finset (set E)} (hF : F.nonempty) (hAF : ∀ B ∈ F, is_exposed 𝕜 A B) : is_exposed 𝕜 A (⋂₀ F) := begin @@ -164,41 +184,17 @@ begin exact hC.inter_left hCA, end -protected lemma is_extreme (hAB : is_exposed 𝕜 A B) : - is_extreme 𝕜 A B := -begin - refine ⟨hAB.subset, λ x₁ hx₁A x₂ hx₂A x hxB hx, _⟩, - obtain ⟨l, rfl⟩ := hAB ⟨x, hxB⟩, - have hl : convex_on 𝕜 univ l := l.to_linear_map.convex_on convex_univ, - have hlx₁ := hxB.2 x₁ hx₁A, - have hlx₂ := hxB.2 x₂ hx₂A, - refine ⟨⟨hx₁A, λ y hy, _⟩, ⟨hx₂A, λ y hy, _⟩⟩, - { rw hlx₁.antisymm (hl.le_left_of_right_le (mem_univ _) (mem_univ _) hx hlx₂), - exact hxB.2 y hy }, - { rw hlx₂.antisymm (hl.le_right_of_left_le (mem_univ _) (mem_univ _) hx hlx₁), - exact hxB.2 y hy } -end - -protected lemma convex (hAB : is_exposed 𝕜 A B) (hA : convex 𝕜 A) : - convex 𝕜 B := +protected lemma is_closed [order_closed_topology 𝕜] {A B : set E} + (hAB : is_exposed 𝕜 A B) (hA : is_closed A) : is_closed B := begin obtain rfl | hB := B.eq_empty_or_nonempty, - { exact convex_empty }, - obtain ⟨l, rfl⟩ := hAB hB, - exact λ x₁ hx₁ x₂ hx₂ a b ha hb hab, ⟨hA hx₁.1 hx₂.1 ha hb hab, λ y hy, - ((l.to_linear_map.concave_on convex_univ).convex_ge _ - ⟨mem_univ _, hx₁.2 y hy⟩ ⟨mem_univ _, hx₂.2 y hy⟩ ha hb hab).2⟩, -end - -protected lemma is_closed [order_closed_topology 𝕜] (hAB : is_exposed 𝕜 A B) (hA : is_closed A) : - is_closed B := -begin - obtain ⟨l, a, rfl⟩ := hAB.eq_inter_halfspace, + { simp }, + obtain ⟨l, a, rfl⟩ := hAB.eq_inter_halfspace' hB, exact hA.is_closed_le continuous_on_const l.continuous.continuous_on, end -protected lemma is_compact [order_closed_topology 𝕜] [t2_space E] (hAB : is_exposed 𝕜 A B) - (hA : is_compact A) : is_compact B := +protected lemma is_compact [order_closed_topology 𝕜] [t2_space E] {A B : set E} + (hAB : is_exposed 𝕜 A B) (hA : is_compact A) : is_compact B := is_compact_of_is_closed_subset hA (hAB.is_closed hA.is_closed) hAB.subset end is_exposed @@ -237,7 +233,48 @@ begin exact ⟨hl.1.1, l, λ y hy, ⟨hl.1.2 y hy, λ hxy, hl.2 y ⟨hy, λ z hz, (hl.1.2 z hz).trans hxy⟩⟩⟩, end +end ordered_ring + +section linear_ordered_ring + +variables {𝕜 : Type*} {E : Type*} [topological_space 𝕜] [linear_ordered_ring 𝕜] + [add_comm_monoid E] [topological_space E] [module 𝕜 E] + {A B C : set E} + +namespace is_exposed + +protected lemma convex (hAB : is_exposed 𝕜 A B) (hA : convex 𝕜 A) : + convex 𝕜 B := +begin + obtain rfl | hB := B.eq_empty_or_nonempty, + { exact convex_empty }, + obtain ⟨l, rfl⟩ := hAB hB, + exact λ x₁ hx₁ x₂ hx₂ a b ha hb hab, ⟨hA hx₁.1 hx₂.1 ha hb hab, λ y hy, + ((l.to_linear_map.concave_on convex_univ).convex_ge _ + ⟨mem_univ _, hx₁.2 y hy⟩ ⟨mem_univ _, hx₂.2 y hy⟩ ha hb hab).2⟩, +end + +protected lemma is_extreme (hAB : is_exposed 𝕜 A B) : + is_extreme 𝕜 A B := +begin + refine ⟨hAB.subset, λ x₁ hx₁A x₂ hx₂A x hxB hx, _⟩, + obtain ⟨l, rfl⟩ := hAB ⟨x, hxB⟩, + have hl : convex_on 𝕜 univ l := l.to_linear_map.convex_on convex_univ, + have hlx₁ := hxB.2 x₁ hx₁A, + have hlx₂ := hxB.2 x₂ hx₂A, + refine ⟨⟨hx₁A, λ y hy, _⟩, ⟨hx₂A, λ y hy, _⟩⟩, + { have := @convex_on.le_left_of_right_le 𝕜 E 𝕜 _ _ _, + rw hlx₁.antisymm (hl.le_left_of_right_le (mem_univ _) (mem_univ _) hx hlx₂), + exact hxB.2 y hy }, + { rw hlx₂.antisymm (hl.le_right_of_left_le (mem_univ _) (mem_univ _) hx hlx₁), + exact hxB.2 y hy } +end + +end is_exposed + lemma exposed_points_subset_extreme_points : A.exposed_points 𝕜 ⊆ A.extreme_points 𝕜 := λ x hx, mem_extreme_points_iff_extreme_singleton.2 (mem_exposed_points_iff_exposed_singleton.1 hx).is_extreme + +end linear_ordered_ring From b4f03bc5115df3bf09e78d52eff3b633aa11c138 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 17 Jan 2023 08:54:07 +0000 Subject: [PATCH 0267/1291] fix(analysis/inner_product_space/basic): restore `bilin_form_of_real_inner` (#18093) I was using this downstream as `bilin_form_of_real_inner.to_quadratic_form`, and I don't see a clear replacement. Until we have api connecting sesquilinear forms with quadratic forms, we should not remove `bilin_form` API. This partially reverts #15780. --- src/analysis/inner_product_space/basic.lean | 10 ++++++++++ 1 file changed, 10 insertions(+) diff --git a/src/analysis/inner_product_space/basic.lean b/src/analysis/inner_product_space/basic.lean index fdac512a03076..2ac8ac71ba5b1 100644 --- a/src/analysis/inner_product_space/basic.lean +++ b/src/analysis/inner_product_space/basic.lean @@ -8,6 +8,7 @@ import analysis.complex.basic import analysis.convex.uniform import analysis.normed_space.completion import analysis.normed_space.bounded_linear_maps +import linear_algebra.bilinear_form /-! # Inner product space @@ -437,6 +438,15 @@ linear_map.mk₂'ₛₗ (ring_hom.id 𝕜) (star_ring_end _) (λ x y z, inner_add_left) (λ r x y, inner_smul_left) +/-- The real inner product as a bilinear form. -/ +@[simps] +def bilin_form_of_real_inner : bilin_form ℝ F := +{ bilin := inner, + bilin_add_left := λ x y z, inner_add_left, + bilin_smul_left := λ a x y, inner_smul_left, + bilin_add_right := λ x y z, inner_add_right, + bilin_smul_right := λ a x y, inner_smul_right } + /-- An inner product with a sum on the left. -/ lemma sum_inner {ι : Type*} (s : finset ι) (f : ι → E) (x : E) : ⟪∑ i in s, f i, x⟫ = ∑ i in s, ⟪f i, x⟫ := (sesq_form_of_inner x).map_sum From 01ad394a11bf06b950232720cf7e8fc6b22f0d6a Mon Sep 17 00:00:00 2001 From: Anatole Dedecker Date: Tue, 17 Jan 2023 10:21:03 +0000 Subject: [PATCH 0268/1291] feat(topology/algebra/equicontinuity): a family of group homomorphisms is equicontinuous iff it is equicontinuous at `1` (#17128) --- src/topology/algebra/equicontinuity.lean | 44 ++++++++++++++++++++++++ 1 file changed, 44 insertions(+) create mode 100644 src/topology/algebra/equicontinuity.lean diff --git a/src/topology/algebra/equicontinuity.lean b/src/topology/algebra/equicontinuity.lean new file mode 100644 index 0000000000000..c7fdcc25e2f3c --- /dev/null +++ b/src/topology/algebra/equicontinuity.lean @@ -0,0 +1,44 @@ +/- +Copyright (c) 2022 Anatole Dedecker. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Anatole Dedecker +-/ +import topology.algebra.uniform_convergence + +/-! +# Algebra-related equicontinuity criterions +-/ + +open function +open_locale uniform_convergence + +@[to_additive] lemma equicontinuous_of_equicontinuous_at_one {ι G M hom : Type*} + [topological_space G] [uniform_space M] [group G] [group M] [topological_group G] + [uniform_group M] [monoid_hom_class hom G M] (F : ι → hom) + (hf : equicontinuous_at (coe_fn ∘ F) (1 : G)) : + equicontinuous (coe_fn ∘ F) := +begin + letI : has_coe_to_fun hom (λ _, G → M) := fun_like.has_coe_to_fun, + rw equicontinuous_iff_continuous, + rw equicontinuous_at_iff_continuous_at at hf, + let φ : G →* (ι → M) := + { to_fun := swap (coe_fn ∘ F), + map_one' := by ext; exact map_one _, + map_mul' := λ a b, by ext; exact map_mul _ _ _ }, + exact continuous_of_continuous_at_one φ hf +end + +@[to_additive] lemma uniform_equicontinuous_of_equicontinuous_at_one {ι G M hom : Type*} + [uniform_space G] [uniform_space M] [group G] [group M] [uniform_group G] [uniform_group M] + [monoid_hom_class hom G M] (F : ι → hom) (hf : equicontinuous_at (coe_fn ∘ F) (1 : G)) : + uniform_equicontinuous (coe_fn ∘ F) := +begin + letI : has_coe_to_fun hom (λ _, G → M) := fun_like.has_coe_to_fun, + rw uniform_equicontinuous_iff_uniform_continuous, + rw equicontinuous_at_iff_continuous_at at hf, + let φ : G →* (ι → M) := + { to_fun := swap (coe_fn ∘ F), + map_one' := by ext; exact map_one _, + map_mul' := λ a b, by ext; exact map_mul _ _ _ }, + exact uniform_continuous_of_continuous_at_one φ hf +end From c46f021fa72ede9249bf50e8068279ff4cfc1f2f Mon Sep 17 00:00:00 2001 From: Anatole Dedecker Date: Tue, 17 Jan 2023 10:21:06 +0000 Subject: [PATCH 0269/1291] feat(analysis/seminorm): characterize continuity of seminorms by closed balls (#17249) --- .../locally_convex/with_seminorms.lean | 43 +++---- src/analysis/seminorm.lean | 111 +++++++++++++----- 2 files changed, 105 insertions(+), 49 deletions(-) diff --git a/src/analysis/locally_convex/with_seminorms.lean b/src/analysis/locally_convex/with_seminorms.lean index c53496a47375a..163db96ce222c 100644 --- a/src/analysis/locally_convex/with_seminorms.lean +++ b/src/analysis/locally_convex/with_seminorms.lean @@ -49,7 +49,7 @@ seminorm, locally convex open normed_field set seminorm topological_space open_locale big_operators nnreal pointwise topological_space -variables {𝕜 𝕜₂ E F G ι ι' : Type*} +variables {𝕜 𝕜₂ 𝕝 𝕝₂ E F G ι ι' : Type*} section filter_basis @@ -404,11 +404,11 @@ begin exact add_group_filter_basis.nhds_zero_eq _, end -lemma with_seminorms.continuous_seminorm [module ℝ E] [normed_algebra ℝ 𝕜] [is_scalar_tower ℝ 𝕜 E] - [has_continuous_const_smul ℝ E] {p : seminorm_family 𝕜 E ι} (hp : with_seminorms p) +lemma with_seminorms.continuous_seminorm [nontrivially_normed_field 𝕝] + [module 𝕝 E] [has_continuous_const_smul 𝕝 E] {p : seminorm_family 𝕝 E ι} (hp : with_seminorms p) (i : ι) : continuous (p i) := begin - refine seminorm.continuous _, + refine seminorm.continuous one_pos _, rw [p.with_seminorms_iff_nhds_eq_infi.mp hp, ball_zero_eq_preimage_ball], exact filter.mem_infi_of_mem i (filter.preimage_mem_comap $ metric.ball_mem_nhds _ one_pos) end @@ -539,15 +539,18 @@ section continuous_bounded namespace seminorm -variables [normed_field 𝕜] [add_comm_group E] [module 𝕜 E] -variables [normed_field 𝕜₂] [add_comm_group F] [module 𝕜₂ F] +variables [nontrivially_normed_field 𝕜] [add_comm_group E] [module 𝕜 E] +variables [normed_field 𝕝] [module 𝕝 E] +variables [nontrivially_normed_field 𝕜₂] [add_comm_group F] [module 𝕜₂ F] +variables [normed_field 𝕝₂] [module 𝕝₂ F] variables {σ₁₂ : 𝕜 →+* 𝕜₂} [ring_hom_isometric σ₁₂] +variables {τ₁₂ : 𝕝 →+* 𝕝₂} [ring_hom_isometric τ₁₂] variables [nonempty ι] [nonempty ι'] -lemma continuous_of_continuous_comp {q : seminorm_family 𝕜₂ F ι'} +lemma continuous_of_continuous_comp {q : seminorm_family 𝕝₂ F ι'} [topological_space E] [topological_add_group E] [topological_space F] [topological_add_group F] (hq : with_seminorms q) - (f : E →ₛₗ[σ₁₂] F) (hf : ∀ i, continuous ((q i).comp f)) : continuous f := + (f : E →ₛₗ[τ₁₂] F) (hf : ∀ i, continuous ((q i).comp f)) : continuous f := begin refine continuous_of_continuous_at_zero f _, simp_rw [continuous_at, f.map_zero, q.with_seminorms_iff_nhds_eq_infi.mp hq, filter.tendsto_infi, @@ -557,17 +560,17 @@ begin exact (map_zero _).symm end -lemma continuous_iff_continuous_comp [normed_algebra ℝ 𝕜₂] [module ℝ F] [is_scalar_tower ℝ 𝕜₂ F] +lemma continuous_iff_continuous_comp {q : seminorm_family 𝕜₂ F ι'} [topological_space E] [topological_add_group E] - [topological_space F] [topological_add_group F] [has_continuous_const_smul ℝ F] + [topological_space F] [topological_add_group F] [has_continuous_const_smul 𝕜₂ F] (hq : with_seminorms q) (f : E →ₛₗ[σ₁₂] F) : continuous f ↔ ∀ i, continuous ((q i).comp f) := ⟨λ h i, continuous.comp (hq.continuous_seminorm i) h, continuous_of_continuous_comp hq f⟩ -lemma continuous_from_bounded {p : seminorm_family 𝕜 E ι} {q : seminorm_family 𝕜₂ F ι'} +lemma continuous_from_bounded {p : seminorm_family 𝕝 E ι} {q : seminorm_family 𝕝₂ F ι'} [topological_space E] [topological_add_group E] (hp : with_seminorms p) [topological_space F] [topological_add_group F] (hq : with_seminorms q) - (f : E →ₛₗ[σ₁₂] F) (hf : seminorm.is_bounded p q f) : continuous f := + (f : E →ₛₗ[τ₁₂] F) (hf : seminorm.is_bounded p q f) : continuous f := begin refine continuous_of_continuous_comp hq _ (λ i, seminorm.continuous_of_continuous_at_zero _), rw [metric.continuous_at_iff', map_zero], @@ -582,23 +585,23 @@ begin refl end -lemma cont_with_seminorms_normed_space (F) [seminormed_add_comm_group F] [normed_space 𝕜₂ F] +lemma cont_with_seminorms_normed_space (F) [seminormed_add_comm_group F] [normed_space 𝕝₂ F] [uniform_space E] [uniform_add_group E] - {p : ι → seminorm 𝕜 E} (hp : with_seminorms p) (f : E →ₛₗ[σ₁₂] F) - (hf : ∃ (s : finset ι) C : ℝ≥0, C ≠ 0 ∧ (norm_seminorm 𝕜₂ F).comp f ≤ C • s.sup p) : + {p : ι → seminorm 𝕝 E} (hp : with_seminorms p) (f : E →ₛₗ[τ₁₂] F) + (hf : ∃ (s : finset ι) C : ℝ≥0, C ≠ 0 ∧ (norm_seminorm 𝕝₂ F).comp f ≤ C • s.sup p) : continuous f := begin rw ←seminorm.is_bounded_const (fin 1) at hf, - exact continuous_from_bounded hp (norm_with_seminorms 𝕜₂ F) f hf, + exact continuous_from_bounded hp (norm_with_seminorms 𝕝₂ F) f hf, end -lemma cont_normed_space_to_with_seminorms (E) [seminormed_add_comm_group E] [normed_space 𝕜 E] +lemma cont_normed_space_to_with_seminorms (E) [seminormed_add_comm_group E] [normed_space 𝕝 E] [uniform_space F] [uniform_add_group F] - {q : ι → seminorm 𝕜₂ F} (hq : with_seminorms q) (f : E →ₛₗ[σ₁₂] F) - (hf : ∀ i : ι, ∃ C : ℝ≥0, C ≠ 0 ∧ (q i).comp f ≤ C • (norm_seminorm 𝕜 E)) : continuous f := + {q : ι → seminorm 𝕝₂ F} (hq : with_seminorms q) (f : E →ₛₗ[τ₁₂] F) + (hf : ∀ i : ι, ∃ C : ℝ≥0, C ≠ 0 ∧ (q i).comp f ≤ C • (norm_seminorm 𝕝 E)) : continuous f := begin rw ←seminorm.const_is_bounded (fin 1) at hf, - exact continuous_from_bounded (norm_with_seminorms 𝕜 E) hq f hf, + exact continuous_from_bounded (norm_with_seminorms 𝕝 E) hq f hf, end end seminorm diff --git a/src/analysis/seminorm.lean b/src/analysis/seminorm.lean index e939177b3737f..deba1e470de3f 100644 --- a/src/analysis/seminorm.lean +++ b/src/analysis/seminorm.lean @@ -38,7 +38,7 @@ set_option old_structure_cmd true open normed_field set open_locale big_operators nnreal pointwise topological_space -variables {R R' 𝕜 𝕜₂ 𝕜₃ E E₂ E₃ F G ι : Type*} +variables {R R' 𝕜 𝕜₂ 𝕜₃ 𝕝 E E₂ E₃ F G ι : Type*} /-- A seminorm on a module over a normed ring is a function to the reals that is positive semidefinite, positive homogeneous, and subadditive. -/ @@ -780,6 +780,21 @@ section normed_field variables [normed_field 𝕜] [add_comm_group E] [module 𝕜 E] (p : seminorm 𝕜 E) {A B : set E} {a : 𝕜} {r : ℝ} {x : E} +lemma ball_norm_mul_subset {p : seminorm 𝕜 E} {k : 𝕜} {r : ℝ} : + p.ball 0 (‖k‖ * r) ⊆ k • p.ball 0 r := +begin + rcases eq_or_ne k 0 with (rfl | hk), + { rw [norm_zero, zero_mul, ball_eq_emptyset _ le_rfl], + exact empty_subset _ }, + { intro x, + rw [set.mem_smul_set, seminorm.mem_ball_zero], + refine λ hx, ⟨k⁻¹ • x, _, _⟩, + { rwa [seminorm.mem_ball_zero, map_smul_eq_mul, norm_inv, + ←(mul_lt_mul_left $ norm_pos_iff.mpr hk), ←mul_assoc, ←(div_eq_mul_inv ‖k‖ ‖k‖), + div_self (ne_of_gt $ norm_pos_iff.mpr hk), one_mul] }, + rw [←smul_assoc, smul_eq_mul, ←div_eq_mul_inv, div_self hk, one_smul] } +end + lemma smul_ball_zero {p : seminorm 𝕜 E} {k : 𝕜} {r : ℝ} (hk : k ≠ 0) : k • p.ball 0 r = p.ball 0 (‖k‖ * r) := begin @@ -788,6 +803,27 @@ begin norm_inv, ← div_eq_inv_mul, div_lt_iff (norm_pos_iff.2 hk), mul_comm] end +lemma smul_closed_ball_subset {p : seminorm 𝕜 E} {k : 𝕜} {r : ℝ} : + k • p.closed_ball 0 r ⊆ p.closed_ball 0 (‖k‖ * r) := +begin + rintros x ⟨y, hy, h⟩, + rw [seminorm.mem_closed_ball_zero, ←h, map_smul_eq_mul], + rw seminorm.mem_closed_ball_zero at hy, + exact mul_le_mul_of_nonneg_left hy (norm_nonneg _) +end + +lemma smul_closed_ball_zero {p : seminorm 𝕜 E} {k : 𝕜} {r : ℝ} (hk : 0 < ‖k‖) : + k • p.closed_ball 0 r = p.closed_ball 0 (‖k‖ * r) := +begin + refine subset_antisymm smul_closed_ball_subset _, + intro x, + rw [set.mem_smul_set, seminorm.mem_closed_ball_zero], + refine λ hx, ⟨k⁻¹ • x, _, _⟩, + { rwa [seminorm.mem_closed_ball_zero, map_smul_eq_mul, norm_inv, ←(mul_le_mul_left hk), + ←mul_assoc, ←(div_eq_mul_inv ‖k‖ ‖k‖), div_self (ne_of_gt hk), one_mul] }, + rw [←smul_assoc, smul_eq_mul, ←div_eq_mul_inv, div_self (norm_pos_iff.mp hk), one_smul], +end + lemma ball_zero_absorbs_ball_zero (p : seminorm 𝕜 E) {r₁ r₂ : ℝ} (hr₁ : 0 < r₁) : absorbs 𝕜 (p.ball 0 r₁) (p.ball 0 r₂) := begin @@ -901,33 +937,43 @@ rfl (p.restrict_scalars 𝕜).ball = p.ball := rfl +@[simp] lemma restrict_scalars_closed_ball (p : seminorm 𝕜' E) : + (p.restrict_scalars 𝕜).closed_ball = p.closed_ball := +rfl + end restrict_scalars /-! ### Continuity criterions for seminorms -/ section continuity -variables [semi_normed_ring 𝕜] [add_comm_group E] - [module 𝕜 E] +variables [nontrivially_normed_field 𝕜] [semi_normed_ring 𝕝] [add_comm_group E] [module 𝕜 E] +variables [module 𝕝 E] -lemma continuous_at_zero [norm_one_class 𝕜] [normed_algebra ℝ 𝕜] [module ℝ E] - [is_scalar_tower ℝ 𝕜 E] [topological_space E] [has_continuous_const_smul ℝ E] {p : seminorm 𝕜 E} - (hp : p.ball 0 1 ∈ (𝓝 0 : filter E)) : +lemma continuous_at_zero' [topological_space E] [has_continuous_const_smul 𝕜 E] {p : seminorm 𝕜 E} + {r : ℝ} (hr : 0 < r) (hp : p.closed_ball 0 r ∈ (𝓝 0 : filter E)) : continuous_at p 0 := begin - change continuous_at (p.restrict_scalars ℝ) 0, - rw ← p.restrict_scalars_ball ℝ at hp, - refine metric.nhds_basis_ball.tendsto_right_iff.mpr _, + refine metric.nhds_basis_closed_ball.tendsto_right_iff.mpr _, intros ε hε, rw map_zero, - suffices : (p.restrict_scalars ℝ).ball 0 ε ∈ (𝓝 0 : filter E), - { rwa seminorm.ball_zero_eq_preimage_ball at this }, - have := (set_smul_mem_nhds_zero_iff hε.ne.symm).mpr hp, - rwa [seminorm.smul_ball_zero hε.ne', real.norm_of_nonneg hε.le, mul_one] at this + suffices : p.closed_ball 0 ε ∈ (𝓝 0 : filter E), + { rwa seminorm.closed_ball_zero_eq_preimage_closed_ball at this }, + rcases exists_norm_lt 𝕜 (div_pos hε hr) with ⟨k, hk0, hkε⟩, + have hk0' := norm_pos_iff.mp hk0, + have := (set_smul_mem_nhds_zero_iff hk0').mpr hp, + refine filter.mem_of_superset this (smul_set_subset_iff.mpr $ λ x hx, _), + rw [mem_closed_ball_zero, map_smul_eq_mul, ← div_mul_cancel ε hr.ne.symm], + exact mul_le_mul hkε.le (p.mem_closed_ball_zero.mp hx) (map_nonneg _ _) (div_nonneg hε.le hr.le) end +lemma continuous_at_zero [topological_space E] [has_continuous_const_smul 𝕜 E] {p : seminorm 𝕜 E} + {r : ℝ} (hr : 0 < r) (hp : p.ball 0 r ∈ (𝓝 0 : filter E)) : + continuous_at p 0 := +continuous_at_zero' hr (filter.mem_of_superset hp $ p.ball_subset_closed_ball _ _) + protected lemma uniform_continuous_of_continuous_at_zero [uniform_space E] [uniform_add_group E] - {p : seminorm 𝕜 E} (hp : continuous_at p 0) : + {p : seminorm 𝕝 E} (hp : continuous_at p 0) : uniform_continuous p := begin have hp : filter.tendsto p (𝓝 0) (𝓝 0) := map_zero p ▸ hp, @@ -938,7 +984,7 @@ begin end protected lemma continuous_of_continuous_at_zero [topological_space E] [topological_add_group E] - {p : seminorm 𝕜 E} (hp : continuous_at p 0) : + {p : seminorm 𝕝 E} (hp : continuous_at p 0) : continuous p := begin letI := topological_add_group.to_uniform_space E, @@ -946,24 +992,31 @@ begin exact (seminorm.uniform_continuous_of_continuous_at_zero hp).continuous end -protected lemma uniform_continuous [norm_one_class 𝕜] [normed_algebra ℝ 𝕜] [module ℝ E] - [is_scalar_tower ℝ 𝕜 E] [uniform_space E] [uniform_add_group E] [has_continuous_const_smul ℝ E] - {p : seminorm 𝕜 E} (hp : p.ball 0 1 ∈ (𝓝 0 : filter E)) : - uniform_continuous p := -seminorm.uniform_continuous_of_continuous_at_zero (continuous_at_zero hp) +protected lemma uniform_continuous [uniform_space E] [uniform_add_group E] + [has_continuous_const_smul 𝕜 E] {p : seminorm 𝕜 E} {r : ℝ} (hr : 0 < r) + (hp : p.ball 0 r ∈ (𝓝 0 : filter E)) : uniform_continuous p := +seminorm.uniform_continuous_of_continuous_at_zero (continuous_at_zero hr hp) -protected lemma continuous [norm_one_class 𝕜] [normed_algebra ℝ 𝕜] [module ℝ E] - [is_scalar_tower ℝ 𝕜 E] [topological_space E] [topological_add_group E] - [has_continuous_const_smul ℝ E] {p : seminorm 𝕜 E} (hp : p.ball 0 1 ∈ (𝓝 0 : filter E)) : - continuous p := -seminorm.continuous_of_continuous_at_zero (continuous_at_zero hp) +protected lemma uniform_continuous' [uniform_space E] [uniform_add_group E] + [has_continuous_const_smul 𝕜 E] {p : seminorm 𝕜 E} {r : ℝ} (hr : 0 < r) + (hp : p.closed_ball 0 r ∈ (𝓝 0 : filter E)) : uniform_continuous p := +seminorm.uniform_continuous_of_continuous_at_zero (continuous_at_zero' hr hp) + +protected lemma continuous [topological_space E] [topological_add_group E] + [has_continuous_const_smul 𝕜 E] {p : seminorm 𝕜 E} {r : ℝ} (hr : 0 < r) + (hp : p.ball 0 r ∈ (𝓝 0 : filter E)) : continuous p := +seminorm.continuous_of_continuous_at_zero (continuous_at_zero hr hp) + +protected lemma continuous' [topological_space E] [topological_add_group E] + [has_continuous_const_smul 𝕜 E] {p : seminorm 𝕜 E} {r : ℝ} (hr : 0 < r) + (hp : p.closed_ball 0 r ∈ (𝓝 0 : filter E)) : continuous p := +seminorm.continuous_of_continuous_at_zero (continuous_at_zero' hr hp) -lemma continuous_of_le [norm_one_class 𝕜] [normed_algebra ℝ 𝕜] [module ℝ E] - [is_scalar_tower ℝ 𝕜 E] [topological_space E] [topological_add_group E] - [has_continuous_const_smul ℝ E] {p q : seminorm 𝕜 E} (hq : continuous q) (hpq : p ≤ q) : +lemma continuous_of_le [topological_space E] [topological_add_group E] + [has_continuous_const_smul 𝕜 E] {p q : seminorm 𝕜 E} (hq : continuous q) (hpq : p ≤ q) : continuous p := begin - refine seminorm.continuous (filter.mem_of_superset + refine seminorm.continuous one_pos (filter.mem_of_superset (is_open.mem_nhds _ $ q.mem_ball_self zero_lt_one) (ball_antitone hpq)), rw ball_zero_eq, exact is_open_lt hq continuous_const From 18562995238888382f71dba0187ac487ae167b98 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 17 Jan 2023 10:21:07 +0000 Subject: [PATCH 0270/1291] feat(group_theory/perm/cycle/basic): Countable sets admit cycles (#17909) If `s` is countable, then there exists a permutation whose support is contained in `s` and which is a cycle on `s`. --- src/group_theory/perm/cycle/basic.lean | 68 ++++++++++++++++++++++++++ 1 file changed, 68 insertions(+) diff --git a/src/group_theory/perm/cycle/basic.lean b/src/group_theory/perm/cycle/basic.lean index 29a35fa756831..8341e96b2727e 100644 --- a/src/group_theory/perm/cycle/basic.lean +++ b/src/group_theory/perm/cycle/basic.lean @@ -7,6 +7,7 @@ import algebra.module.big_operators import data.finset.noncomm_prod import data.fintype.perm import data.int.modeq +import group_theory.perm.list import group_theory.perm.sign import logic.equiv.fintype /-! @@ -790,6 +791,13 @@ lemma is_cycle_on.extend_domain {p : β → Prop} [decidable_pred p] (f : α ≃ (g.extend_domain f).is_cycle_on (coe ∘ f '' s) := ⟨h.1.extend_domain, by { rintro _ ⟨a, ha, rfl⟩ _ ⟨b, hb, rfl⟩, exact (h.2 ha hb).extend_domain }⟩ +protected lemma is_cycle_on.countable (hs : f.is_cycle_on s) : s.countable := +begin + obtain rfl | ⟨a, ha⟩ := s.eq_empty_or_nonempty, + { exact set.countable_empty }, + { exact (set.countable_range $ λ n : ℤ, (⇑(f ^ n) : α → α) a).mono (hs.2 ha) } +end + end is_cycle_on /-! @@ -1614,9 +1622,69 @@ end equiv.perm open equiv +namespace list +variables [decidable_eq α] {l : list α} + +lemma nodup.is_cycle_on_form_perm (h : l.nodup) : l.form_perm.is_cycle_on {a | a ∈ l} := +begin + refine ⟨l.form_perm.bij_on (λ _, form_perm_mem_iff_mem), λ a ha b hb, _⟩, + rw [set.mem_set_of, ←index_of_lt_length] at ha hb, + rw [←index_of_nth_le ha, ←index_of_nth_le hb], + refine ⟨l.index_of b - l.index_of a, _⟩, + simp only [sub_eq_neg_add, zpow_add, zpow_neg, equiv.perm.inv_eq_iff_eq, zpow_coe_nat, + equiv.perm.coe_mul, form_perm_pow_apply_nth_le _ h], + rw add_comm, +end + +end list + +namespace int +open equiv + +lemma add_left_one_is_cycle : (equiv.add_left 1 : perm ℤ).is_cycle := +⟨0, one_ne_zero, λ n _, ⟨n, by simp⟩⟩ + +lemma add_right_one_is_cycle : (equiv.add_right 1 : perm ℤ).is_cycle := +⟨0, one_ne_zero, λ n _, ⟨n, by simp⟩⟩ + +end int + +namespace finset +variables [decidable_eq α] [fintype α] + +lemma exists_cycle_on (s : finset α) : ∃ f : perm α, f.is_cycle_on s ∧ f.support ⊆ s := +begin + refine ⟨s.to_list.form_perm, _, + λ x hx, by simpa using list.mem_of_form_perm_apply_ne _ _ (perm.mem_support.1 hx)⟩, + convert s.nodup_to_list.is_cycle_on_form_perm, + simp, +end + +end finset + namespace set variables {f : perm α} {s : set α} +lemma countable.exists_cycle_on (hs : s.countable) : + ∃ f : perm α, f.is_cycle_on s ∧ {x | f x ≠ x} ⊆ s := +begin + classical, + obtain hs' | hs' := s.finite_or_infinite, + { refine ⟨hs'.to_finset.to_list.form_perm, _, + λ x hx, by simpa using list.mem_of_form_perm_apply_ne _ _ hx⟩, + convert hs'.to_finset.nodup_to_list.is_cycle_on_form_perm, + simp }, + haveI := hs.to_subtype, + haveI := hs'.to_subtype, + obtain ⟨f⟩ : nonempty (ℤ ≃ s) := infer_instance, + refine ⟨(equiv.add_right 1).extend_domain f, _, λ x hx, of_not_not $ λ h, hx $ + perm.extend_domain_apply_not_subtype _ _ h⟩, + convert int.add_right_one_is_cycle.is_cycle_on.extend_domain _, + rw [image_comp, equiv.image_eq_preimage], + ext, + simp, +end + lemma prod_self_eq_Union_perm (hf : f.is_cycle_on s) : s ×ˢ s = ⋃ n : ℤ, (λ a, (a, (f ^ n) a)) '' s := begin From a993e7ee843ff31f2d474197483d2f6d6610addd Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 17 Jan 2023 10:21:08 +0000 Subject: [PATCH 0271/1291] feat(algebra/order/to_interval_mod): more relations between `to_Ixx_{mod/div}` definitions (#17933) `to_Ico_div` and `to_Ioc_div` differ by at most 1, and the modulo versions differ by at most `b`. This also renames some lemmas that have the wrong name. --- src/algebra/order/to_interval_mod.lean | 60 +++++++++++++++++++++++--- src/topology/instances/add_circle.lean | 2 +- 2 files changed, 56 insertions(+), 6 deletions(-) diff --git a/src/algebra/order/to_interval_mod.lean b/src/algebra/order/to_interval_mod.lean index 88b61cd9a7006..f696f550e5552 100644 --- a/src/algebra/order/to_interval_mod.lean +++ b/src/algebra/order/to_interval_mod.lean @@ -6,6 +6,7 @@ Authors: Joseph Myers import algebra.module.basic import algebra.order.archimedean import algebra.periodic +import data.int.succ_pred import group_theory.quotient_group /-! @@ -454,6 +455,8 @@ begin rw [hz, to_Ioc_mod_zsmul_add] } end +/-! ### Links between the `Ico` and `Ioc` variants applied to the same element -/ + section Ico_Ioc variables (a : α) {b : α} (hb : 0 < b) (x : α) @@ -496,6 +499,14 @@ lemma mem_Ioo_mod_iff_to_Ico_mod_add_period_ne_to_Ioc_mod : mem_Ioo_mod a b x ↔ to_Ico_mod a hb x + b ≠ to_Ioc_mod a hb x := (tfae_mem_Ioo_mod a hb x).out 0 2 lemma mem_Ioo_mod_iff_to_Ico_mod_ne_left : mem_Ioo_mod a b x ↔ to_Ico_mod a hb x ≠ a := (tfae_mem_Ioo_mod a hb x).out 0 3 + +lemma not_mem_Ioo_mod_iff_to_Ico_mod_add_period_eq_to_Ioc_mod : + ¬mem_Ioo_mod a b x ↔ to_Ico_mod a hb x + b = to_Ioc_mod a hb x := +(mem_Ioo_mod_iff_to_Ico_mod_add_period_ne_to_Ioc_mod hb).not_left + +lemma not_mem_Ioo_mod_iff_to_Ico_mod_eq_left : ¬mem_Ioo_mod a b x ↔ to_Ico_mod a hb x = a := +(mem_Ioo_mod_iff_to_Ico_mod_ne_left hb).not_left + lemma mem_Ioo_mod_iff_to_Ioc_mod_ne_right : mem_Ioo_mod a b x ↔ to_Ioc_mod a hb x ≠ a + b := begin rw [mem_Ioo_mod_iff_to_Ico_mod_eq_to_Ioc_mod, to_Ico_mod_eq_iff hb], @@ -504,20 +515,27 @@ begin (to_Ioc_mod_add_to_Ioc_div_zsmul _ _ _).symm⟩⟩, end +lemma not_mem_Ioo_mod_iff_to_Ioc_eq_right : ¬mem_Ioo_mod a b x ↔ to_Ioc_mod a hb x = a + b := +(mem_Ioo_mod_iff_to_Ioc_mod_ne_right hb).not_left + lemma mem_Ioo_mod_iff_to_Ico_div_eq_to_Ioc_div : mem_Ioo_mod a b x ↔ to_Ico_div a hb x = to_Ioc_div a hb x := by rw [mem_Ioo_mod_iff_to_Ico_mod_eq_to_Ioc_mod hb, to_Ico_mod, to_Ioc_mod, sub_right_inj, (zsmul_strict_mono_left hb).injective.eq_iff] -lemma mem_Ioo_mod_iff_to_Ico_div_add_one_ne_to_Ioc_div : +lemma mem_Ioo_mod_iff_to_Ico_div_ne_to_Ioc_div_add_one : mem_Ioo_mod a b x ↔ to_Ico_div a hb x ≠ to_Ioc_div a hb x + 1 := by rw [mem_Ioo_mod_iff_to_Ico_mod_add_period_ne_to_Ioc_mod hb, ne, ne, to_Ico_mod, to_Ioc_mod, ← eq_sub_iff_add_eq, sub_sub, sub_right_inj, ← add_one_zsmul, (zsmul_strict_mono_left hb).injective.eq_iff] +lemma not_mem_Ioo_mod_iff_to_Ico_div_eq_to_Ioc_div_add_one : + ¬mem_Ioo_mod a b x ↔ to_Ico_div a hb x = to_Ioc_div a hb x + 1 := +(mem_Ioo_mod_iff_to_Ico_div_ne_to_Ioc_div_add_one hb).not_left + include hb -lemma mem_Ioo_mod_iff_sub_ne_zsmul : mem_Ioo_mod a b x ↔ ∀ z : ℤ, x ≠ a + z • b := +lemma mem_Ioo_mod_iff_ne_add_zsmul : mem_Ioo_mod a b x ↔ ∀ z : ℤ, x ≠ a + z • b := begin rw [mem_Ioo_mod_iff_to_Ico_mod_ne_left hb, ← not_iff_not], push_neg, split; intro h, @@ -527,10 +545,17 @@ begin exact ⟨lt_add_of_pos_right a hb, h⟩, }, end -lemma mem_Ioo_mod_iff_eq_mod_zmultiples : +lemma not_mem_Ioo_mod_iff_eq_add_zsmul : ¬mem_Ioo_mod a b x ↔ ∃ z : ℤ, x = a + z • b := +by simpa only [not_forall, not_ne_iff] using (mem_Ioo_mod_iff_ne_add_zsmul hb).not + +lemma not_mem_Ioo_mod_iff_eq_mod_zmultiples : + ¬mem_Ioo_mod a b x ↔ (x : α ⧸ add_subgroup.zmultiples b) = a := +by simp_rw [not_mem_Ioo_mod_iff_eq_add_zsmul hb, quotient_add_group.eq_iff_sub_mem, + add_subgroup.mem_zmultiples_iff, eq_sub_iff_add_eq', eq_comm] + +lemma mem_Ioo_mod_iff_ne_mod_zmultiples : mem_Ioo_mod a b x ↔ (x : α ⧸ add_subgroup.zmultiples b) ≠ a := -by simp_rw [mem_Ioo_mod_iff_sub_ne_zsmul hb, ne, quotient_add_group.eq_iff_sub_mem, - add_subgroup.mem_zmultiples_iff, eq_sub_iff_add_eq', eq_comm, not_exists] +(not_mem_Ioo_mod_iff_eq_mod_zmultiples hb).not_right lemma Ico_eq_locus_Ioc_eq_Union_Ioo : {x | to_Ico_mod a hb x = to_Ioc_mod a hb x} = ⋃ z : ℤ, set.Ioo (a + z • b) (a + b + z • b) := @@ -539,6 +564,31 @@ begin exact (mem_Ioo_mod_iff_to_Ico_mod_eq_to_Ioc_mod hb).symm, end +lemma to_Ioc_div_wcovby_to_Ico_div (a : α) {b : α} (hb : 0 < b) (x : α) : + to_Ioc_div a hb x ⩿ to_Ico_div a hb x := +begin + suffices : to_Ioc_div a hb x = to_Ico_div a hb x ∨ to_Ioc_div a hb x + 1 = to_Ico_div a hb x, + { rwa [wcovby_iff_eq_or_covby, ←order.succ_eq_iff_covby] }, + rw [eq_comm, ←mem_Ioo_mod_iff_to_Ico_div_eq_to_Ioc_div, + eq_comm, ←not_mem_Ioo_mod_iff_to_Ico_div_eq_to_Ioc_div_add_one], + exact em _, +end + +lemma to_Ico_mod_le_to_Ioc_mod (a : α) {b : α} (hb : 0 < b) (x : α) : + to_Ico_mod a hb x ≤ to_Ioc_mod a hb x := +begin + rw [to_Ico_mod, to_Ioc_mod, sub_le_sub_iff_left], + exact zsmul_mono_left hb.le (to_Ioc_div_wcovby_to_Ico_div _ _ _).le +end + +lemma to_Ioc_mod_le_to_Ico_mod_add (a : α) {b : α} (hb : 0 < b) (x : α) : + to_Ioc_mod a hb x ≤ to_Ico_mod a hb x + b := +begin + rw [to_Ico_mod, to_Ioc_mod, sub_add, sub_le_sub_iff_left, sub_le_iff_le_add, ←add_one_zsmul, + (zsmul_strict_mono_left hb).le_iff_le], + apply (to_Ioc_div_wcovby_to_Ico_div _ _ _).le_succ, +end + end Ico_Ioc lemma to_Ico_mod_eq_self {a b x : α} (hb : 0 < b) : to_Ico_mod a hb x = x ↔ x ∈ set.Ico a (a + b) := diff --git a/src/topology/instances/add_circle.lean b/src/topology/instances/add_circle.lean index 98f0214adc573..371831392e71f 100644 --- a/src/topology/instances/add_circle.lean +++ b/src/topology/instances/add_circle.lean @@ -89,7 +89,7 @@ variables {x} (hx : (x : 𝕜 ⧸ zmultiples p) ≠ a) lemma to_Ico_mod_eventually_eq_to_Ioc_mod : to_Ico_mod a hp =ᶠ[𝓝 x] to_Ioc_mod a hp := is_open.mem_nhds (by {rw Ico_eq_locus_Ioc_eq_Union_Ioo, exact is_open_Union (λ i, is_open_Ioo)}) $ - (mem_Ioo_mod_iff_to_Ico_mod_eq_to_Ioc_mod hp).1 ((mem_Ioo_mod_iff_eq_mod_zmultiples hp).2 hx) + (mem_Ioo_mod_iff_to_Ico_mod_eq_to_Ioc_mod hp).1 ((mem_Ioo_mod_iff_ne_mod_zmultiples hp).2 hx) lemma continuous_at_to_Ico_mod : continuous_at (to_Ico_mod a hp) x := let h := to_Ico_mod_eventually_eq_to_Ioc_mod a hp hx in continuous_at_iff_continuous_left_right.2 $ From 9956c3806d0f9553e5c6e6af68970563a1619cd1 Mon Sep 17 00:00:00 2001 From: Michael Stoll Date: Tue, 17 Jan 2023 13:26:00 +0000 Subject: [PATCH 0272/1291] feat(number_theory/diophantine_approximation): add more results (#18193) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR adds the following results: ```lean lemma rat.finite_rat_abs_sub_lt_one_div_denom_sq (ξ : ℚ) : {q : ℚ | |ξ - q| < 1 / q.denom ^ 2}.finite lemma real.infinite_rat_abs_sub_lt_one_div_denom_sq_iff_irrational (ξ : ℝ) : {q : ℚ | |ξ - q| < 1 / q.denom ^ 2}.infinite ↔ irrational ξ ``` See the [discussion on Zulip](https://leanprover.zulipchat.com/#narrow/stream/116395-maths/topic/Diophantine.20approximation/near/321399454). --- .../diophantine_approximation.lean | 92 ++++++++++++++++++- 1 file changed, 89 insertions(+), 3 deletions(-) diff --git a/src/number_theory/diophantine_approximation.lean b/src/number_theory/diophantine_approximation.lean index 61334a8a3aa0b..b3bf3e84d7cce 100644 --- a/src/number_theory/diophantine_approximation.lean +++ b/src/number_theory/diophantine_approximation.lean @@ -27,12 +27,20 @@ The main results are three variants of Dirichlet's approximation theorem: satisfying `|ξ - q| ≤ 1/((n+1)*q.denom)` and `q.denom ≤ n`, and -* `dioph_approx.rat_approx_infinite`, which states that for irrational `ξ`, the set - `{q : ℚ | |ξ - q| < 1/q.denom^2}` is infinite, +* `real.infinite_rat_abs_sub_lt_one_div_denom_sq_of_irrational`, which states that + for irrational `ξ`, the set `{q : ℚ | |ξ - q| < 1/q.denom^2}` is infinite. + +We also show a converse, +* `rat.finite_rat_abs_sub_lt_one_div_denom_sq`, which states that the set above is finite + when `ξ` is a rational number. + +Both statements are combined to give an equivalence, +`real.infinite_rat_abs_sub_lt_one_div_denom_sq_iff_irrational`. ## Implementation notes -We use the namespace `real` for the results. +We use the namespace `real` for the results on real numbers and `rat` for the results +on rational numbers. ## References @@ -179,3 +187,81 @@ end end rat_approx end real + +namespace rat + +/-! +### Finitely many good approximations to rational numbers + +We now show that a rational number `ξ` has only finitely many good rational +approximations. +-/ + +open set + +/-- If `ξ` is rational, then the good rational approximations to `ξ` have bounded +numerator and denominator. -/ +lemma denom_le_and_le_num_le_of_sub_lt_one_div_denom_sq {ξ q : ℚ} (h : |ξ - q| < 1 / q.denom ^ 2) : + q.denom ≤ ξ.denom ∧ ⌈ξ * q.denom⌉ - 1 ≤ q.num ∧ q.num ≤ ⌊ξ * q.denom⌋ + 1 := +begin + have hq₀ : (0 : ℚ) < q.denom := nat.cast_pos.mpr q.pos, + replace h : |ξ * q.denom - q.num| < 1 / q.denom, + { rw ← mul_lt_mul_right hq₀ at h, + conv_lhs at h { rw [← abs_of_pos hq₀, ← abs_mul, sub_mul, mul_denom_eq_num], }, + rwa [sq, div_mul, mul_div_cancel_left _ hq₀.ne'] at h, }, + split, + { rcases eq_or_ne ξ q with rfl | H, + { exact le_rfl, }, + { have hξ₀ : (0 : ℚ) < ξ.denom := nat.cast_pos.mpr ξ.pos, + rw [← rat.num_div_denom ξ, div_mul_eq_mul_div, div_sub' _ _ _ hξ₀.ne', abs_div, + abs_of_pos hξ₀, div_lt_iff hξ₀, div_mul_comm, mul_one] at h, + refine nat.cast_le.mp (((one_lt_div hq₀).mp $ lt_of_le_of_lt _ h).le), + norm_cast, + rw [mul_comm _ q.num], + exact int.one_le_abs (sub_ne_zero_of_ne $ mt rat.eq_iff_mul_eq_mul.mpr H), } }, + { obtain ⟨h₁, h₂⟩ := abs_sub_lt_iff.mp (h.trans_le $ (one_div_le zero_lt_one hq₀).mp $ + (@one_div_one ℚ _).symm ▸ nat.cast_le.mpr q.pos), + rw [sub_lt_iff_lt_add, add_comm] at h₁ h₂, + rw [← sub_lt_iff_lt_add] at h₂, + norm_cast at h₁ h₂, + exact ⟨sub_le_iff_le_add.mpr (int.ceil_le.mpr h₁.le), + sub_le_iff_le_add.mp (int.le_floor.mpr h₂.le)⟩, } +end + +/-- A rational number has only finitely many good rational approximations. -/ +lemma finite_rat_abs_sub_lt_one_div_denom_sq (ξ : ℚ) : + {q : ℚ | |ξ - q| < 1 / q.denom ^ 2}.finite := +begin + let f : ℚ → ℤ × ℕ := λ q, (q.num, q.denom), + set s := {q : ℚ | |ξ - q| < 1 / q.denom ^ 2}, + have hinj : function.injective f, + { intros a b hab, + simp only [prod.mk.inj_iff] at hab, + rw [← rat.num_div_denom a, ← rat.num_div_denom b, hab.1, hab.2], }, + have H : f '' s ⊆ ⋃ (y : ℕ) (hy : y ∈ Ioc 0 ξ.denom), Icc (⌈ξ * y⌉ - 1) (⌊ξ * y⌋ + 1) ×ˢ {y}, + { intros xy hxy, + simp only [mem_image, mem_set_of_eq] at hxy, + obtain ⟨q, hq₁, hq₂⟩ := hxy, + obtain ⟨hd, hn⟩ := denom_le_and_le_num_le_of_sub_lt_one_div_denom_sq hq₁, + simp_rw [mem_Union], + refine ⟨q.denom, set.mem_Ioc.mpr ⟨q.pos, hd⟩, _⟩, + simp only [prod_singleton, mem_image, mem_Icc, (congr_arg prod.snd (eq.symm hq₂)).trans rfl], + exact ⟨q.num, hn, hq₂⟩, }, + refine finite.of_finite_image (finite.subset _ H) (inj_on_of_injective hinj s), + exact finite.bUnion (finite_Ioc _ _) (λ x hx, finite.prod (finite_Icc _ _) (finite_singleton _)), +end + +end rat + +/-- The set of good rational approximations to a real number `ξ` is infinite if and only if +`ξ` is irrational. -/ +lemma real.infinite_rat_abs_sub_lt_one_div_denom_sq_iff_irrational (ξ : ℝ) : + {q : ℚ | |ξ - q| < 1 / q.denom ^ 2}.infinite ↔ irrational ξ := +begin + refine ⟨λ h, (irrational_iff_ne_rational ξ).mpr (λ a b H, set.not_infinite.mpr _ h), + real.infinite_rat_abs_sub_lt_one_div_denom_sq_of_irrational⟩, + convert rat.finite_rat_abs_sub_lt_one_div_denom_sq ((a : ℚ) / b), + ext q, + rw [H, (by push_cast : (1 : ℝ) / q.denom ^ 2 = (1 / q.denom ^ 2 : ℚ))], + norm_cast, +end From e04043d6bf7264a3c84bc69711dc354958ca4516 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Tue, 17 Jan 2023 15:24:34 +0000 Subject: [PATCH 0273/1291] chore(*): add mathlib4 synchronization comments (#18196) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.gcd_monoid.multiset` * `algebra.hom.group_action` * `algebra.order.smul` * `data.finset.basic` * `data.finset.card` * `data.finset.fin` * `data.finset.fold` * `data.finset.image` * `data.finset.option` * `data.finset.order` * `data.finset.pi` * `data.finset.prod` * `data.finset.sum` * `data.fintype.basic` * `data.fintype.pi` * `data.int.succ_pred` * `data.multiset.antidiagonal` * `data.set.constructions` * `data.set.intervals.iso_Ioo` --- src/algebra/gcd_monoid/multiset.lean | 3 +++ src/algebra/hom/group_action.lean | 3 +++ src/algebra/order/smul.lean | 3 +++ src/data/finset/basic.lean | 3 +++ src/data/finset/card.lean | 3 +++ src/data/finset/fin.lean | 3 +++ src/data/finset/fold.lean | 3 +++ src/data/finset/image.lean | 3 +++ src/data/finset/option.lean | 3 +++ src/data/finset/order.lean | 3 +++ src/data/finset/pi.lean | 3 +++ src/data/finset/prod.lean | 3 +++ src/data/finset/sum.lean | 3 +++ src/data/fintype/basic.lean | 3 +++ src/data/fintype/pi.lean | 3 +++ src/data/int/succ_pred.lean | 3 +++ src/data/multiset/antidiagonal.lean | 3 +++ src/data/set/constructions.lean | 3 +++ src/data/set/intervals/iso_Ioo.lean | 3 +++ 19 files changed, 57 insertions(+) diff --git a/src/algebra/gcd_monoid/multiset.lean b/src/algebra/gcd_monoid/multiset.lean index 1647017c603c6..bfe56ae028968 100644 --- a/src/algebra/gcd_monoid/multiset.lean +++ b/src/algebra/gcd_monoid/multiset.lean @@ -10,6 +10,9 @@ import data.multiset.fold /-! # GCD and LCM operations on multisets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions - `multiset.gcd` - the greatest common denominator of a `multiset` of elements of a `gcd_monoid` diff --git a/src/algebra/hom/group_action.lean b/src/algebra/hom/group_action.lean index a2e4343592e30..8c086115d579e 100644 --- a/src/algebra/hom/group_action.lean +++ b/src/algebra/hom/group_action.lean @@ -9,6 +9,9 @@ import algebra.module.basic /-! # Equivariant homomorphisms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions * `mul_action_hom M X Y`, the type of equivariant functions from `X` to `Y`, where `M` is a monoid diff --git a/src/algebra/order/smul.lean b/src/algebra/order/smul.lean index 9082ac44abeca..b6cd7435b978b 100644 --- a/src/algebra/order/smul.lean +++ b/src/algebra/order/smul.lean @@ -13,6 +13,9 @@ import tactic.positivity /-! # Ordered scalar product +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define * `ordered_smul R M` : an ordered additive commutative monoid `M` is an `ordered_smul` diff --git a/src/data/finset/basic.lean b/src/data/finset/basic.lean index 69d110ee1a825..8e804a6b8e8ae 100644 --- a/src/data/finset/basic.lean +++ b/src/data/finset/basic.lean @@ -11,6 +11,9 @@ import tactic.monotonicity /-! # Finite sets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Terms of type `finset α` are one way of talking about finite subsets of `α` in mathlib. Below, `finset α` is defined as a structure with 2 fields: diff --git a/src/data/finset/card.lean b/src/data/finset/card.lean index 8a6f1f2e0e381..98cf0399b0e58 100644 --- a/src/data/finset/card.lean +++ b/src/data/finset/card.lean @@ -9,6 +9,9 @@ import tactic.by_contra /-! # Cardinality of a finite set +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This defines the cardinality of a `finset` and provides induction principles for finsets. ## Main declarations diff --git a/src/data/finset/fin.lean b/src/data/finset/fin.lean index 2c5e4de5df438..e034777f9a1b1 100644 --- a/src/data/finset/fin.lean +++ b/src/data/finset/fin.lean @@ -8,6 +8,9 @@ import data.finset.card /-! # Finsets in `fin n` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A few constructions for finsets in `fin n`. ## Main declarations diff --git a/src/data/finset/fold.lean b/src/data/finset/fold.lean index b42a0a14e09a4..9ebaefe51cdc6 100644 --- a/src/data/finset/fold.lean +++ b/src/data/finset/fold.lean @@ -9,6 +9,9 @@ import data.multiset.fold /-! # The fold operation for a commutative associative operation over a finset. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ namespace finset diff --git a/src/data/finset/image.lean b/src/data/finset/image.lean index 601b3bbd93e24..31afc2348852a 100644 --- a/src/data/finset/image.lean +++ b/src/data/finset/image.lean @@ -10,6 +10,9 @@ import data.int.order.basic /-! # Image and map operations on finite sets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Thie file provides the finite analog of `set.image`, along with some other similar functions. Note there are two ways to take the image over a finset; via `finset.image` which applies the diff --git a/src/data/finset/option.lean b/src/data/finset/option.lean index 62acd32e05217..d63ce497a8d5a 100644 --- a/src/data/finset/option.lean +++ b/src/data/finset/option.lean @@ -9,6 +9,9 @@ import order.hom.basic /-! # Finite sets in `option α` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define * `option.to_finset`: construct an empty or singleton `finset α` from an `option α`; diff --git a/src/data/finset/order.lean b/src/data/finset/order.lean index af8b632cd9c4f..188ddc4c16205 100644 --- a/src/data/finset/order.lean +++ b/src/data/finset/order.lean @@ -8,6 +8,9 @@ import data.finset.basic /-! # Finsets of ordered types + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ universes u v w diff --git a/src/data/finset/pi.lean b/src/data/finset/pi.lean index aad4696c2c647..5bab75fe063e1 100644 --- a/src/data/finset/pi.lean +++ b/src/data/finset/pi.lean @@ -8,6 +8,9 @@ import data.multiset.pi /-! # The cartesian product of finsets + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ namespace finset diff --git a/src/data/finset/prod.lean b/src/data/finset/prod.lean index 9bca2f43adcf7..278f9ac72858b 100644 --- a/src/data/finset/prod.lean +++ b/src/data/finset/prod.lean @@ -8,6 +8,9 @@ import data.finset.card /-! # Finsets in product types +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines finset constructions on the product type `α × β`. Beware not to confuse with the `finset.prod` operation which computes the multiplicative product. diff --git a/src/data/finset/sum.lean b/src/data/finset/sum.lean index 8bf824c2d95c4..86d4f7ab8974f 100644 --- a/src/data/finset/sum.lean +++ b/src/data/finset/sum.lean @@ -9,6 +9,9 @@ import data.finset.card /-! # Disjoint sum of finsets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the disjoint sum of two finsets as `finset (α ⊕ β)`. Beware not to confuse with the `finset.sum` operation which computes the additive sum. diff --git a/src/data/fintype/basic.lean b/src/data/fintype/basic.lean index 3c71e4fdf1e89..5603a0ecd5c40 100644 --- a/src/data/fintype/basic.lean +++ b/src/data/fintype/basic.lean @@ -8,6 +8,9 @@ import data.finset.image /-! # Finite types +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines a typeclass to state that a type is finite. ## Main declarations diff --git a/src/data/fintype/pi.lean b/src/data/fintype/pi.lean index cee402273dd3e..2c5a26cf37d81 100644 --- a/src/data/fintype/pi.lean +++ b/src/data/fintype/pi.lean @@ -8,6 +8,9 @@ import data.finset.pi /-! # fintype instances for pi types + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ variables {α : Type*} diff --git a/src/data/int/succ_pred.lean b/src/data/int/succ_pred.lean index 0c68e9c46f7d8..1d65823913dbf 100644 --- a/src/data/int/succ_pred.lean +++ b/src/data/int/succ_pred.lean @@ -9,6 +9,9 @@ import data.nat.succ_pred /-! # Successors and predecessors of integers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we show that `ℤ` is both an archimedean `succ_order` and an archimedean `pred_order`. -/ diff --git a/src/data/multiset/antidiagonal.lean b/src/data/multiset/antidiagonal.lean index 1efc8e414afbd..3a376862c720f 100644 --- a/src/data/multiset/antidiagonal.lean +++ b/src/data/multiset/antidiagonal.lean @@ -8,6 +8,9 @@ import data.multiset.powerset /-! # The antidiagonal on a multiset. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The antidiagonal of a multiset `s` consists of all pairs `(t₁, t₂)` such that `t₁ + t₂ = s`. These pairs are counted with multiplicities. -/ diff --git a/src/data/set/constructions.lean b/src/data/set/constructions.lean index cc49bc5c308b5..2c46bacbb68d9 100644 --- a/src/data/set/constructions.lean +++ b/src/data/set/constructions.lean @@ -8,6 +8,9 @@ import data.finset.basic /-! # Constructions involving sets of sets. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Finite Intersections We define a structure `has_finite_inter` which asserts that a set `S` of subsets of `α` is diff --git a/src/data/set/intervals/iso_Ioo.lean b/src/data/set/intervals/iso_Ioo.lean index a3cfa15ccf236..1ce9692e865e7 100644 --- a/src/data/set/intervals/iso_Ioo.lean +++ b/src/data/set/intervals/iso_Ioo.lean @@ -9,6 +9,9 @@ import tactic.field_simp /-! # Order isomorphism between a linear ordered field and `(-1, 1)` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we provide an order isomorphism `order_iso_Ioo_neg_one_one` between the open interval `(-1, 1)` in a linear ordered field and the whole field. -/ From 6e70e0d419bf686784937d64ed4bfde866ff229e Mon Sep 17 00:00:00 2001 From: Sidharth Hariharan Date: Tue, 17 Jan 2023 18:29:04 +0000 Subject: [PATCH 0274/1291] feat(data/polynomial): Add Partial Fraction Decomposition (Existence) (#17709) Formalisation and proof of the existence of Partial Fraction Decomposition over an integral domain. Co-authored-by: Kevin Buzzard Co-authored-by: Sidharth Hariharan --- src/data/polynomial/partial_fractions.lean | 130 +++++++++++++++++++++ src/ring_theory/integral_domain.lean | 23 ++++ 2 files changed, 153 insertions(+) create mode 100644 src/data/polynomial/partial_fractions.lean diff --git a/src/data/polynomial/partial_fractions.lean b/src/data/polynomial/partial_fractions.lean new file mode 100644 index 0000000000000..8192f1d5bc55d --- /dev/null +++ b/src/data/polynomial/partial_fractions.lean @@ -0,0 +1,130 @@ +/- +Copyright (c) Sidharth Hariharan. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Kevin Buzzard, Sidharth Hariharan +-/ +import data.polynomial.div +import data.zmod.basic +import logic.function.basic +import ring_theory.localization.fraction_ring +import tactic.field_simp +import tactic.linear_combination + +/-! + +# Partial fractions + +These results were formalised by the Xena Project, at the suggestion +of Patrick Massot. + + +# The main theorem + +* `div_eq_quo_add_sum_rem_div`: General partial fraction decomposition theorem for polynomials over + an integral domain R : + If f, g₁, g₂, ..., gₙ ∈ R[X] and the gᵢs are all monic and pairwise coprime, then ∃ q, r₁, ..., rₙ + ∈ R[X] such that f / g₁g₂...gₙ = q + r₁/g₁ + ... + rₙ/gₙ and for all i, deg(rᵢ) < deg(gᵢ). + +* The result is formalized here in slightly more generality, using finsets. That is, if ι is an + arbitrary index type, g denotes a map from ι to R[X], and if s is an arbitrary finite subset of ι, + with g i monic for all i ∈ s and for all i,j ∈ s, i ≠ j → g i is coprime to g j, then we have + ∃ q ∈ R[X] , r : ι → R[X] such that ∀ i ∈ s, deg(r i) < deg(g i) and + f / ∏ g i = q + ∑ (r i) / (g i), where the product and sum are over s. + +* The proof is done by proving the two-denominator case and then performing finset induction for an + arbitrary (finite) number of denominators. + +## Scope for Expansion + +* Proving uniqueness of the decomposition + +-/ + + +variables (R : Type) [comm_ring R] [is_domain R] + +open_locale polynomial + +open polynomial + +variables (K : Type) [field K] [algebra R[X] K] [is_fraction_ring R[X] K] + +section two_denominators + +/-- +Let R be an integral domain and f, g₁, g₂ ∈ R[X]. Let g₁ and g₂ be monic and coprime. +Then, ∃ q, r₁, r₂ ∈ R[X] such that f / g₁g₂ = q + r₁/g₁ + r₂/g₂ and deg(r₁) < deg(g₁) and +deg(r₂) < deg(g₂). +-/ +lemma div_eq_quo_add_rem_div_add_rem_div (f : R[X]) {g₁ g₂ : R[X]} + (hg₁ : g₁.monic) (hg₂ : g₂.monic) (hcoprime : is_coprime g₁ g₂) : + ∃ q r₁ r₂ : R[X], r₁.degree < g₁.degree ∧ r₂.degree < g₂.degree ∧ + (↑f : K) / (↑g₁ * ↑g₂) = ↑q + ↑r₁ / ↑g₁ + ↑r₂ / ↑g₂ := +begin + rcases hcoprime with ⟨c, d, hcd⟩, + refine ⟨(f * d) /ₘ g₁ + (f * c) /ₘ g₂, (f * d) %ₘ g₁, (f * c) %ₘ g₂, + (degree_mod_by_monic_lt _ hg₁), (degree_mod_by_monic_lt _ hg₂), _⟩, + have hg₁' : (↑g₁ : K) ≠ 0, + { norm_cast, exact hg₁.ne_zero_of_ne zero_ne_one, }, + have hg₂' : (↑g₂ : K) ≠ 0, + { norm_cast, exact hg₂.ne_zero_of_ne zero_ne_one, }, + have hfc := mod_by_monic_add_div (f * c) hg₂, + have hfd := mod_by_monic_add_div (f * d) hg₁, + field_simp, + norm_cast, + linear_combination (-1) * f * hcd + (-1) * g₁ * hfc + (-1) * g₂ * hfd, +end + +end two_denominators + +section n_denominators + +open_locale big_operators classical + +/-- +Let R be an integral domain and f ∈ R[X]. Let s be a finite index set. +Then, a fraction of the form f / ∏ (g i) can be rewritten as q + ∑ (r i) / (g i), where +deg(r i) < deg(g i), provided that the g i are monic and pairwise coprime. +-/ +lemma div_eq_quo_add_sum_rem_div (f : R[X]) {ι : Type*} {g : ι → R[X]} {s : finset ι} + (hg : ∀ i ∈ s, (g i).monic) + (hcop : set.pairwise ↑s (λ i j, is_coprime (g i) (g j))) : + ∃ (q : R[X]) (r : ι → R[X]), (∀ i ∈ s, (r i).degree < (g i).degree) ∧ + (↑f : K) / ∏ i in s, ↑(g i) = ↑q + ∑ i in s, ↑(r i) / ↑(g i) := +begin + induction s using finset.induction_on with a b hab Hind f generalizing f, + { refine ⟨f, (λ (i : ι), (0 : R[X])), λ i, _, by simp⟩, + rintro ⟨⟩, }, + obtain ⟨q₀, r₁, r₂, hdeg₁, hdeg₂, (hf : (↑f : K) / _ = _)⟩ := + div_eq_quo_add_rem_div_add_rem_div R K f + (_ : monic (g a)) + (_ : monic ∏ (i : ι) in b, (g i)) + _, + { obtain ⟨q, r, hrdeg, IH⟩ := Hind (λ i hi, hg i (finset.mem_insert_of_mem hi)) + (set.pairwise.mono ( finset.coe_subset.2 $ λ i hi, finset.mem_insert_of_mem hi) hcop) r₂, + refine ⟨q₀ + q, λ i, if i = a then r₁ else r i, _, _⟩, + { intro i, + split_ifs with h1, + { cases h1, + intro _, + exact hdeg₁, }, + { intro hi, + exact hrdeg i (finset.mem_of_mem_insert_of_ne hi h1), }, }, + norm_cast at ⊢ hf IH, + rw [finset.prod_insert hab, hf, IH, finset.sum_insert hab, if_pos rfl], + transitivity (↑(q₀ + q : R[X]) : K) + (↑r₁ / ↑(g a) + ∑ (i : ι) in b, ↑(r i) / ↑(g i)), + { push_cast, ring, }, + congr' 2, + refine finset.sum_congr rfl (λ x hxb, _), + rw if_neg, + rintro rfl, + exact hab hxb }, + { exact hg a (b.mem_insert_self a), }, + { exact monic_prod_of_monic _ _ (λ i hi, hg i (finset.mem_insert_of_mem hi)), }, + { refine is_coprime.prod_right (λ i hi, hcop (finset.mem_coe.2 (b.mem_insert_self a)) + (finset.mem_coe.2 (finset.mem_insert_of_mem hi)) _), + rintro rfl, + exact hab hi, }, +end + +end n_denominators diff --git a/src/ring_theory/integral_domain.lean b/src/ring_theory/integral_domain.lean index 210b69b717ac6..ff73d6a5b1cc9 100644 --- a/src/ring_theory/integral_domain.lean +++ b/src/ring_theory/integral_domain.lean @@ -156,6 +156,29 @@ end end +section euclidean_division + +namespace polynomial + +open_locale polynomial + +variables (K : Type) [field K] [algebra R[X] K] [is_fraction_ring R[X] K] + +lemma div_eq_quo_add_rem_div (f : R[X]) {g : R[X]} (hg : g.monic) : + ∃ q r : R[X], r.degree < g.degree ∧ (↑f : K) / ↑g = ↑q + ↑r / ↑g := +begin + refine ⟨f /ₘ g, f %ₘ g, _, _⟩, + { exact degree_mod_by_monic_lt _ hg }, + { have hg' : (↑g : K) ≠ 0 := by exact_mod_cast (monic.ne_zero hg), + field_simp [hg'], + norm_cast, + rw [add_comm, mul_comm, mod_by_monic_add_div f hg] }, +end + +end polynomial + +end euclidean_division + variables [fintype G] lemma card_fiber_eq_of_mem_range {H : Type*} [group H] [decidable_eq H] From 008205aa645b3f194c1da47025c5f110c8406eab Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Tue, 17 Jan 2023 18:29:05 +0000 Subject: [PATCH 0275/1291] chore(analysis/normed/field/basic): split (#18195) Split material on infinite sums out of `analysis/normed/field/basic` into a new file `/infinite_sum`. This mimics the structure of `analysis/normed/group/*`, and reduces the imports of the former. --- src/analysis/asymptotics/asymptotics.lean | 1 + src/analysis/normed/field/basic.lean | 132 +------------------- src/analysis/normed/field/infinite_sum.lean | 124 ++++++++++++++++++ src/analysis/normed_space/basic.lean | 9 ++ src/analysis/normed_space/exponential.lean | 1 + 5 files changed, 136 insertions(+), 131 deletions(-) create mode 100644 src/analysis/normed/field/infinite_sum.lean diff --git a/src/analysis/asymptotics/asymptotics.lean b/src/analysis/asymptotics/asymptotics.lean index d3af26c25320c..edf8db8b19103 100644 --- a/src/analysis/asymptotics/asymptotics.lean +++ b/src/analysis/asymptotics/asymptotics.lean @@ -3,6 +3,7 @@ Copyright (c) 2019 Jeremy Avigad. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Yury Kudryashov -/ +import analysis.normed.group.infinite_sum import analysis.normed_space.basic import topology.algebra.order.liminf_limsup import topology.local_homeomorph diff --git a/src/analysis/normed/field/basic.lean b/src/analysis/normed/field/basic.lean index f610971286317..8b14076599c1b 100644 --- a/src/analysis/normed/field/basic.lean +++ b/src/analysis/normed/field/basic.lean @@ -4,8 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Patrick Massot, Johannes Hölzl -/ import algebra.algebra.subalgebra.basic -import analysis.normed.group.infinite_sum -import topology.algebra.module.basic +import analysis.normed.group.basic import topology.instances.ennreal import topology.instances.rat @@ -640,14 +639,6 @@ by simp [real.to_nnreal_of_nonneg, nnnorm, norm_of_nonneg, hx] lemma nnnorm_mul_to_nnreal (x : ℝ) {y : ℝ} (hy : 0 ≤ y) : ‖x‖₊ * y.to_nnreal = ‖x * y‖₊ := by simp [real.to_nnreal_of_nonneg, nnnorm, norm_of_nonneg, hy] -/-- If `E` is a nontrivial topological module over `ℝ`, then `E` has no isolated points. -This is a particular case of `module.punctured_nhds_ne_bot`. -/ -instance punctured_nhds_module_ne_bot - {E : Type*} [add_comm_group E] [topological_space E] [has_continuous_add E] [nontrivial E] - [module ℝ E] [has_continuous_smul ℝ E] (x : E) : - ne_bot (𝓝[≠] x) := -module.punctured_nhds_ne_bot ℝ E x - end real namespace nnreal @@ -741,127 +732,6 @@ by simpa only [← nnreal.coe_le_coe, nnreal.coe_mul] using norm_zpow_le_mul_nor end -section cauchy_product - -/-! ## Multiplying two infinite sums in a normed ring - -In this section, we prove various results about `(∑' x : ι, f x) * (∑' y : ι', g y)` in a normed -ring. There are similar results proven in `topology/algebra/infinite_sum` (e.g `tsum_mul_tsum`), -but in a normed ring we get summability results which aren't true in general. - -We first establish results about arbitrary index types, `β` and `γ`, and then we specialize to -`β = γ = ℕ` to prove the Cauchy product formula -(see `tsum_mul_tsum_eq_tsum_sum_antidiagonal_of_summable_norm`). - -### Arbitrary index types --/ - -variables {ι' : Type*} [normed_ring α] - -open finset -open_locale classical - -lemma summable.mul_of_nonneg {f : ι → ℝ} {g : ι' → ℝ} - (hf : summable f) (hg : summable g) (hf' : 0 ≤ f) (hg' : 0 ≤ g) : - summable (λ (x : ι × ι'), f x.1 * g x.2) := -let ⟨s, hf⟩ := hf in -let ⟨t, hg⟩ := hg in -suffices this : ∀ u : finset (ι × ι'), ∑ x in u, f x.1 * g x.2 ≤ s*t, - from summable_of_sum_le (λ x, mul_nonneg (hf' _) (hg' _)) this, -assume u, -calc ∑ x in u, f x.1 * g x.2 - ≤ ∑ x in u.image prod.fst ×ˢ u.image prod.snd, f x.1 * g x.2 : - sum_mono_set_of_nonneg (λ x, mul_nonneg (hf' _) (hg' _)) subset_product -... = ∑ x in u.image prod.fst, ∑ y in u.image prod.snd, f x * g y : sum_product -... = ∑ x in u.image prod.fst, f x * ∑ y in u.image prod.snd, g y : - sum_congr rfl (λ x _, mul_sum.symm) -... ≤ ∑ x in u.image prod.fst, f x * t : - sum_le_sum - (λ x _, mul_le_mul_of_nonneg_left (sum_le_has_sum _ (λ _ _, hg' _) hg) (hf' _)) -... = (∑ x in u.image prod.fst, f x) * t : sum_mul.symm -... ≤ s * t : - mul_le_mul_of_nonneg_right (sum_le_has_sum _ (λ _ _, hf' _) hf) (hg.nonneg $ λ _, hg' _) - -lemma summable.mul_norm {f : ι → α} {g : ι' → α} - (hf : summable (λ x, ‖f x‖)) (hg : summable (λ x, ‖g x‖)) : - summable (λ (x : ι × ι'), ‖f x.1 * g x.2‖) := -summable_of_nonneg_of_le (λ x, norm_nonneg (f x.1 * g x.2)) (λ x, norm_mul_le (f x.1) (g x.2)) - (hf.mul_of_nonneg hg (λ x, norm_nonneg $ f x) (λ x, norm_nonneg $ g x) : _) - -lemma summable_mul_of_summable_norm [complete_space α] {f : ι → α} {g : ι' → α} - (hf : summable (λ x, ‖f x‖)) (hg : summable (λ x, ‖g x‖)) : - summable (λ (x : ι × ι'), f x.1 * g x.2) := -summable_of_summable_norm (hf.mul_norm hg) - -/-- Product of two infinites sums indexed by arbitrary types. - See also `tsum_mul_tsum` if `f` and `g` are *not* absolutely summable. -/ -lemma tsum_mul_tsum_of_summable_norm [complete_space α] {f : ι → α} {g : ι' → α} - (hf : summable (λ x, ‖f x‖)) (hg : summable (λ x, ‖g x‖)) : - (∑' x, f x) * (∑' y, g y) = (∑' z : ι × ι', f z.1 * g z.2) := -tsum_mul_tsum (summable_of_summable_norm hf) (summable_of_summable_norm hg) - (summable_mul_of_summable_norm hf hg) - -/-! ### `ℕ`-indexed families (Cauchy product) - -We prove two versions of the Cauchy product formula. The first one is -`tsum_mul_tsum_eq_tsum_sum_range_of_summable_norm`, where the `n`-th term is a sum over -`finset.range (n+1)` involving `nat` subtraction. -In order to avoid `nat` subtraction, we also provide -`tsum_mul_tsum_eq_tsum_sum_antidiagonal_of_summable_norm`, -where the `n`-th term is a sum over all pairs `(k, l)` such that `k+l=n`, which corresponds to the -`finset` `finset.nat.antidiagonal n`. -/ - -section nat - -open finset.nat - -lemma summable_norm_sum_mul_antidiagonal_of_summable_norm {f g : ℕ → α} - (hf : summable (λ x, ‖f x‖)) (hg : summable (λ x, ‖g x‖)) : - summable (λ n, ‖∑ kl in antidiagonal n, f kl.1 * g kl.2‖) := -begin - have := summable_sum_mul_antidiagonal_of_summable_mul - (summable.mul_of_nonneg hf hg (λ _, norm_nonneg _) (λ _, norm_nonneg _)), - refine summable_of_nonneg_of_le (λ _, norm_nonneg _) _ this, - intros n, - calc ‖∑ kl in antidiagonal n, f kl.1 * g kl.2‖ - ≤ ∑ kl in antidiagonal n, ‖f kl.1 * g kl.2‖ : norm_sum_le _ _ - ... ≤ ∑ kl in antidiagonal n, ‖f kl.1‖ * ‖g kl.2‖ : sum_le_sum (λ i _, norm_mul_le _ _) -end - -/-- The Cauchy product formula for the product of two infinite sums indexed by `ℕ`, - expressed by summing on `finset.nat.antidiagonal`. - See also `tsum_mul_tsum_eq_tsum_sum_antidiagonal` if `f` and `g` are - *not* absolutely summable. -/ -lemma tsum_mul_tsum_eq_tsum_sum_antidiagonal_of_summable_norm [complete_space α] {f g : ℕ → α} - (hf : summable (λ x, ‖f x‖)) (hg : summable (λ x, ‖g x‖)) : - (∑' n, f n) * (∑' n, g n) = ∑' n, ∑ kl in antidiagonal n, f kl.1 * g kl.2 := -tsum_mul_tsum_eq_tsum_sum_antidiagonal (summable_of_summable_norm hf) (summable_of_summable_norm hg) - (summable_mul_of_summable_norm hf hg) - -lemma summable_norm_sum_mul_range_of_summable_norm {f g : ℕ → α} - (hf : summable (λ x, ‖f x‖)) (hg : summable (λ x, ‖g x‖)) : - summable (λ n, ‖∑ k in range (n+1), f k * g (n - k)‖) := -begin - simp_rw ← sum_antidiagonal_eq_sum_range_succ (λ k l, f k * g l), - exact summable_norm_sum_mul_antidiagonal_of_summable_norm hf hg -end - -/-- The Cauchy product formula for the product of two infinite sums indexed by `ℕ`, - expressed by summing on `finset.range`. - See also `tsum_mul_tsum_eq_tsum_sum_range` if `f` and `g` are - *not* absolutely summable. -/ -lemma tsum_mul_tsum_eq_tsum_sum_range_of_summable_norm [complete_space α] {f g : ℕ → α} - (hf : summable (λ x, ‖f x‖)) (hg : summable (λ x, ‖g x‖)) : - (∑' n, f n) * (∑' n, g n) = ∑' n, ∑ k in range (n+1), f k * g (n - k) := -begin - simp_rw ← sum_antidiagonal_eq_sum_range_succ (λ k l, f k * g l), - exact tsum_mul_tsum_eq_tsum_sum_antidiagonal_of_summable_norm hf hg -end - -end nat - -end cauchy_product - section ring_hom_isometric variables {R₁ : Type*} {R₂ : Type*} {R₃ : Type*} diff --git a/src/analysis/normed/field/infinite_sum.lean b/src/analysis/normed/field/infinite_sum.lean new file mode 100644 index 0000000000000..eb12b5e6bd2e4 --- /dev/null +++ b/src/analysis/normed/field/infinite_sum.lean @@ -0,0 +1,124 @@ +/- +Copyright (c) 2021 Anatole Dedecker. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Anatole Dedecker +-/ +import analysis.normed.field.basic +import analysis.normed.group.infinite_sum + +/-! # Multiplying two infinite sums in a normed ring + +In this file, we prove various results about `(∑' x : ι, f x) * (∑' y : ι', g y)` in a normed +ring. There are similar results proven in `topology/algebra/infinite_sum` (e.g `tsum_mul_tsum`), +but in a normed ring we get summability results which aren't true in general. + +We first establish results about arbitrary index types, `β` and `γ`, and then we specialize to +`β = γ = ℕ` to prove the Cauchy product formula +(see `tsum_mul_tsum_eq_tsum_sum_antidiagonal_of_summable_norm`). +!-/ + +variables {α : Type*} {ι : Type*} {ι' : Type*} [normed_ring α] + +open_locale big_operators classical +open finset + +/-! ### Arbitrary index types -/ + +lemma summable.mul_of_nonneg {f : ι → ℝ} {g : ι' → ℝ} + (hf : summable f) (hg : summable g) (hf' : 0 ≤ f) (hg' : 0 ≤ g) : + summable (λ (x : ι × ι'), f x.1 * g x.2) := +let ⟨s, hf⟩ := hf in +let ⟨t, hg⟩ := hg in +suffices this : ∀ u : finset (ι × ι'), ∑ x in u, f x.1 * g x.2 ≤ s*t, + from summable_of_sum_le (λ x, mul_nonneg (hf' _) (hg' _)) this, +assume u, +calc ∑ x in u, f x.1 * g x.2 + ≤ ∑ x in u.image prod.fst ×ˢ u.image prod.snd, f x.1 * g x.2 : + sum_mono_set_of_nonneg (λ x, mul_nonneg (hf' _) (hg' _)) subset_product +... = ∑ x in u.image prod.fst, ∑ y in u.image prod.snd, f x * g y : sum_product +... = ∑ x in u.image prod.fst, f x * ∑ y in u.image prod.snd, g y : + sum_congr rfl (λ x _, mul_sum.symm) +... ≤ ∑ x in u.image prod.fst, f x * t : + sum_le_sum + (λ x _, mul_le_mul_of_nonneg_left (sum_le_has_sum _ (λ _ _, hg' _) hg) (hf' _)) +... = (∑ x in u.image prod.fst, f x) * t : sum_mul.symm +... ≤ s * t : + mul_le_mul_of_nonneg_right (sum_le_has_sum _ (λ _ _, hf' _) hf) (hg.nonneg $ λ _, hg' _) + +lemma summable.mul_norm {f : ι → α} {g : ι' → α} + (hf : summable (λ x, ‖f x‖)) (hg : summable (λ x, ‖g x‖)) : + summable (λ (x : ι × ι'), ‖f x.1 * g x.2‖) := +summable_of_nonneg_of_le (λ x, norm_nonneg (f x.1 * g x.2)) (λ x, norm_mul_le (f x.1) (g x.2)) + (hf.mul_of_nonneg hg (λ x, norm_nonneg $ f x) (λ x, norm_nonneg $ g x) : _) + +lemma summable_mul_of_summable_norm [complete_space α] {f : ι → α} {g : ι' → α} + (hf : summable (λ x, ‖f x‖)) (hg : summable (λ x, ‖g x‖)) : + summable (λ (x : ι × ι'), f x.1 * g x.2) := +summable_of_summable_norm (hf.mul_norm hg) + +/-- Product of two infinites sums indexed by arbitrary types. + See also `tsum_mul_tsum` if `f` and `g` are *not* absolutely summable. -/ +lemma tsum_mul_tsum_of_summable_norm [complete_space α] {f : ι → α} {g : ι' → α} + (hf : summable (λ x, ‖f x‖)) (hg : summable (λ x, ‖g x‖)) : + (∑' x, f x) * (∑' y, g y) = (∑' z : ι × ι', f z.1 * g z.2) := +tsum_mul_tsum (summable_of_summable_norm hf) (summable_of_summable_norm hg) + (summable_mul_of_summable_norm hf hg) + +/-! ### `ℕ`-indexed families (Cauchy product) + +We prove two versions of the Cauchy product formula. The first one is +`tsum_mul_tsum_eq_tsum_sum_range_of_summable_norm`, where the `n`-th term is a sum over +`finset.range (n+1)` involving `nat` subtraction. +In order to avoid `nat` subtraction, we also provide +`tsum_mul_tsum_eq_tsum_sum_antidiagonal_of_summable_norm`, +where the `n`-th term is a sum over all pairs `(k, l)` such that `k+l=n`, which corresponds to the +`finset` `finset.nat.antidiagonal n`. -/ + +section nat + +open finset.nat + +lemma summable_norm_sum_mul_antidiagonal_of_summable_norm {f g : ℕ → α} + (hf : summable (λ x, ‖f x‖)) (hg : summable (λ x, ‖g x‖)) : + summable (λ n, ‖∑ kl in antidiagonal n, f kl.1 * g kl.2‖) := +begin + have := summable_sum_mul_antidiagonal_of_summable_mul + (summable.mul_of_nonneg hf hg (λ _, norm_nonneg _) (λ _, norm_nonneg _)), + refine summable_of_nonneg_of_le (λ _, norm_nonneg _) _ this, + intros n, + calc ‖∑ kl in antidiagonal n, f kl.1 * g kl.2‖ + ≤ ∑ kl in antidiagonal n, ‖f kl.1 * g kl.2‖ : norm_sum_le _ _ + ... ≤ ∑ kl in antidiagonal n, ‖f kl.1‖ * ‖g kl.2‖ : sum_le_sum (λ i _, norm_mul_le _ _) +end + +/-- The Cauchy product formula for the product of two infinite sums indexed by `ℕ`, + expressed by summing on `finset.nat.antidiagonal`. + See also `tsum_mul_tsum_eq_tsum_sum_antidiagonal` if `f` and `g` are + *not* absolutely summable. -/ +lemma tsum_mul_tsum_eq_tsum_sum_antidiagonal_of_summable_norm [complete_space α] {f g : ℕ → α} + (hf : summable (λ x, ‖f x‖)) (hg : summable (λ x, ‖g x‖)) : + (∑' n, f n) * (∑' n, g n) = ∑' n, ∑ kl in antidiagonal n, f kl.1 * g kl.2 := +tsum_mul_tsum_eq_tsum_sum_antidiagonal (summable_of_summable_norm hf) (summable_of_summable_norm hg) + (summable_mul_of_summable_norm hf hg) + +lemma summable_norm_sum_mul_range_of_summable_norm {f g : ℕ → α} + (hf : summable (λ x, ‖f x‖)) (hg : summable (λ x, ‖g x‖)) : + summable (λ n, ‖∑ k in range (n+1), f k * g (n - k)‖) := +begin + simp_rw ← sum_antidiagonal_eq_sum_range_succ (λ k l, f k * g l), + exact summable_norm_sum_mul_antidiagonal_of_summable_norm hf hg +end + +/-- The Cauchy product formula for the product of two infinite sums indexed by `ℕ`, + expressed by summing on `finset.range`. + See also `tsum_mul_tsum_eq_tsum_sum_range` if `f` and `g` are + *not* absolutely summable. -/ +lemma tsum_mul_tsum_eq_tsum_sum_range_of_summable_norm [complete_space α] {f g : ℕ → α} + (hf : summable (λ x, ‖f x‖)) (hg : summable (λ x, ‖g x‖)) : + (∑' n, f n) * (∑' n, g n) = ∑' n, ∑ k in range (n+1), f k * g (n - k) := +begin + simp_rw ← sum_antidiagonal_eq_sum_range_succ (λ k l, f k * g l), + exact tsum_mul_tsum_eq_tsum_sum_antidiagonal_of_summable_norm hf hg +end + +end nat diff --git a/src/analysis/normed_space/basic.lean b/src/analysis/normed_space/basic.lean index 1d493c820861b..ca451a24b6b02 100644 --- a/src/analysis/normed_space/basic.lean +++ b/src/analysis/normed_space/basic.lean @@ -7,6 +7,7 @@ import algebra.algebra.pi import algebra.algebra.restrict_scalars import analysis.normed.field.basic import data.real.sqrt +import topology.algebra.module.basic /-! # Normed spaces @@ -341,6 +342,14 @@ lemma nnnorm_surjective : surjective (nnnorm : E → ℝ≥0) := end surj +/-- If `E` is a nontrivial topological module over `ℝ`, then `E` has no isolated points. +This is a particular case of `module.punctured_nhds_ne_bot`. -/ +instance real.punctured_nhds_module_ne_bot + {E : Type*} [add_comm_group E] [topological_space E] [has_continuous_add E] [nontrivial E] + [module ℝ E] [has_continuous_smul ℝ E] (x : E) : + ne_bot (𝓝[≠] x) := +module.punctured_nhds_ne_bot ℝ E x + theorem interior_closed_ball' [normed_space ℝ E] [nontrivial E] (x : E) (r : ℝ) : interior (closed_ball x r) = ball x r := begin diff --git a/src/analysis/normed_space/exponential.lean b/src/analysis/normed_space/exponential.lean index af316d21047c5..86df10c323a9c 100644 --- a/src/analysis/normed_space/exponential.lean +++ b/src/analysis/normed_space/exponential.lean @@ -5,6 +5,7 @@ Authors: Anatole Dedecker, Eric Wieser -/ import analysis.analytic.basic import analysis.complex.basic +import analysis.normed.field.infinite_sum import data.nat.choose.cast import data.finset.noncomm_prod import topology.algebra.algebra From 3249a849dbc131395156322a6fdce9611ef83375 Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Tue, 17 Jan 2023 20:29:28 +0000 Subject: [PATCH 0276/1291] chore(analysis/normed_space/star/basic): split (#18194) The import `analysis/normed_space/operator_norm` was added to this file by @j-loreaux in #16964, but nowadays `operator_norm` is quite a heavy import (it contains everything on the strong topologies, etc), whereas I hope that `normed_space/star/basic` can become a fairly lightweight one (because it is imported by `is_R_or_C` which is imported everywhere). I propose to split out the material from #16964. --- src/analysis/normed_space/star/basic.lean | 57 -------------------- src/analysis/normed_space/star/mul.lean | 64 +++++++++++++++++++++++ 2 files changed, 64 insertions(+), 57 deletions(-) create mode 100644 src/analysis/normed_space/star/mul.lean diff --git a/src/analysis/normed_space/star/basic.lean b/src/analysis/normed_space/star/basic.lean index f504ce482ba0e..6b65310a500d4 100644 --- a/src/analysis/normed_space/star/basic.lean +++ b/src/analysis/normed_space/star/basic.lean @@ -7,7 +7,6 @@ Authors: Frédéric Dupuis import analysis.normed.group.hom import analysis.normed_space.basic import analysis.normed_space.linear_isometry -import analysis.normed_space.operator_norm import algebra.star.self_adjoint import algebra.star.unitary @@ -255,59 +254,3 @@ variables {𝕜} lemma starₗᵢ_apply {x : E} : starₗᵢ 𝕜 x = star x := rfl end starₗᵢ - -section mul - -open continuous_linear_map - -variables (𝕜) [densely_normed_field 𝕜] [non_unital_normed_ring E] [star_ring E] [cstar_ring E] -variables [normed_space 𝕜 E] [is_scalar_tower 𝕜 E E] [smul_comm_class 𝕜 E E] (a : E) - -/-- In a C⋆-algebra `E`, either unital or non-unital, multiplication on the left by `a : E` has -norm equal to the norm of `a`. -/ -@[simp] lemma op_nnnorm_mul : ‖mul 𝕜 E a‖₊ = ‖a‖₊ := -begin - rw ←Sup_closed_unit_ball_eq_nnnorm, - refine cSup_eq_of_forall_le_of_forall_lt_exists_gt _ _ (λ r hr, _), - { exact (metric.nonempty_closed_ball.mpr zero_le_one).image _ }, - { rintro - ⟨x, hx, rfl⟩, - exact ((mul 𝕜 E a).unit_le_op_norm x $ mem_closed_ball_zero_iff.mp hx).trans - (op_norm_mul_apply_le 𝕜 E a) }, - { have ha : 0 < ‖a‖₊ := zero_le'.trans_lt hr, - rw [←inv_inv (‖a‖₊), nnreal.lt_inv_iff_mul_lt (inv_ne_zero ha.ne')] at hr, - obtain ⟨k, hk₁, hk₂⟩ := normed_field.exists_lt_nnnorm_lt 𝕜 (mul_lt_mul_of_pos_right hr $ - nnreal.inv_pos.2 ha), - refine ⟨_, ⟨k • star a, _, rfl⟩, _⟩, - { simpa only [mem_closed_ball_zero_iff, norm_smul, one_mul, norm_star] using - (nnreal.le_inv_iff_mul_le ha.ne').1 (one_mul ‖a‖₊⁻¹ ▸ hk₂.le : ‖k‖₊ ≤ ‖a‖₊⁻¹) }, - { simp only [map_smul, nnnorm_smul, mul_apply', mul_smul_comm, cstar_ring.nnnorm_self_mul_star], - rwa [←nnreal.div_lt_iff (mul_pos ha ha).ne', div_eq_mul_inv, mul_inv, ←mul_assoc] } }, -end - -/-- In a C⋆-algebra `E`, either unital or non-unital, multiplication on the right by `a : E` has -norm eqaul to the norm of `a`. -/ -@[simp] lemma op_nnnorm_mul_flip : ‖(mul 𝕜 E).flip a‖₊ = ‖a‖₊ := -begin - rw [←Sup_unit_ball_eq_nnnorm, ←nnnorm_star, ←@op_nnnorm_mul 𝕜 E, ←Sup_unit_ball_eq_nnnorm], - congr' 1, - simp only [mul_apply', flip_apply], - refine set.subset.antisymm _ _; - rintro - ⟨b, hb, rfl⟩; - refine ⟨star b, by simpa only [norm_star, mem_ball_zero_iff] using hb, _⟩, - { simp only [←star_mul, nnnorm_star] }, - { simpa using (nnnorm_star (star b * a)).symm } -end - -variables (E) - -/-- In a C⋆-algebra `E`, either unital or non-unital, the left regular representation is an -isometry. -/ -lemma mul_isometry : isometry (mul 𝕜 E) := -add_monoid_hom_class.isometry_of_norm _ (λ a, congr_arg coe $ op_nnnorm_mul 𝕜 a) - -/-- In a C⋆-algebra `E`, either unital or non-unital, the right regular anti-representation is an -isometry. -/ -lemma mul_flip_isometry : isometry (mul 𝕜 E).flip := -add_monoid_hom_class.isometry_of_norm _ (λ a, congr_arg coe $ op_nnnorm_mul_flip 𝕜 a) - -end mul diff --git a/src/analysis/normed_space/star/mul.lean b/src/analysis/normed_space/star/mul.lean new file mode 100644 index 0000000000000..55dbf1ba90df0 --- /dev/null +++ b/src/analysis/normed_space/star/mul.lean @@ -0,0 +1,64 @@ +/- +Copyright (c) 2022 Jireh Loreaux. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jireh Loreaux +-/ +import analysis.normed_space.star.basic +import analysis.normed_space.operator_norm + +/-! # The left-regular representation is an isometry for C⋆-algebras -/ + +open continuous_linear_map + +local postfix `⋆`:std.prec.max_plus := star + +variables (𝕜 : Type*) {E : Type*} +variables [densely_normed_field 𝕜] [non_unital_normed_ring E] [star_ring E] [cstar_ring E] +variables [normed_space 𝕜 E] [is_scalar_tower 𝕜 E E] [smul_comm_class 𝕜 E E] (a : E) + +/-- In a C⋆-algebra `E`, either unital or non-unital, multiplication on the left by `a : E` has +norm equal to the norm of `a`. -/ +@[simp] lemma op_nnnorm_mul : ‖mul 𝕜 E a‖₊ = ‖a‖₊ := +begin + rw ←Sup_closed_unit_ball_eq_nnnorm, + refine cSup_eq_of_forall_le_of_forall_lt_exists_gt _ _ (λ r hr, _), + { exact (metric.nonempty_closed_ball.mpr zero_le_one).image _ }, + { rintro - ⟨x, hx, rfl⟩, + exact ((mul 𝕜 E a).unit_le_op_norm x $ mem_closed_ball_zero_iff.mp hx).trans + (op_norm_mul_apply_le 𝕜 E a) }, + { have ha : 0 < ‖a‖₊ := zero_le'.trans_lt hr, + rw [←inv_inv (‖a‖₊), nnreal.lt_inv_iff_mul_lt (inv_ne_zero ha.ne')] at hr, + obtain ⟨k, hk₁, hk₂⟩ := normed_field.exists_lt_nnnorm_lt 𝕜 (mul_lt_mul_of_pos_right hr $ + nnreal.inv_pos.2 ha), + refine ⟨_, ⟨k • star a, _, rfl⟩, _⟩, + { simpa only [mem_closed_ball_zero_iff, norm_smul, one_mul, norm_star] using + (nnreal.le_inv_iff_mul_le ha.ne').1 (one_mul ‖a‖₊⁻¹ ▸ hk₂.le : ‖k‖₊ ≤ ‖a‖₊⁻¹) }, + { simp only [map_smul, nnnorm_smul, mul_apply', mul_smul_comm, cstar_ring.nnnorm_self_mul_star], + rwa [←nnreal.div_lt_iff (mul_pos ha ha).ne', div_eq_mul_inv, mul_inv, ←mul_assoc] } }, +end + +/-- In a C⋆-algebra `E`, either unital or non-unital, multiplication on the right by `a : E` has +norm eqaul to the norm of `a`. -/ +@[simp] lemma op_nnnorm_mul_flip : ‖(mul 𝕜 E).flip a‖₊ = ‖a‖₊ := +begin + rw [←Sup_unit_ball_eq_nnnorm, ←nnnorm_star, ←@op_nnnorm_mul 𝕜 E, ←Sup_unit_ball_eq_nnnorm], + congr' 1, + simp only [mul_apply', flip_apply], + refine set.subset.antisymm _ _; + rintro - ⟨b, hb, rfl⟩; + refine ⟨star b, by simpa only [norm_star, mem_ball_zero_iff] using hb, _⟩, + { simp only [←star_mul, nnnorm_star] }, + { simpa using (nnnorm_star (star b * a)).symm } +end + +variables (E) + +/-- In a C⋆-algebra `E`, either unital or non-unital, the left regular representation is an +isometry. -/ +lemma mul_isometry : isometry (mul 𝕜 E) := +add_monoid_hom_class.isometry_of_norm _ (λ a, congr_arg coe $ op_nnnorm_mul 𝕜 a) + +/-- In a C⋆-algebra `E`, either unital or non-unital, the right regular anti-representation is an +isometry. -/ +lemma mul_flip_isometry : isometry (mul 𝕜 E).flip := +add_monoid_hom_class.isometry_of_norm _ (λ a, congr_arg coe $ op_nnnorm_mul_flip 𝕜 a) From cbdf7b565832144d024caa5a550117c6df0204a5 Mon Sep 17 00:00:00 2001 From: Paul Lezeau <72892199+Paul-Lez@users.noreply.github.com> Date: Tue, 17 Jan 2023 23:27:00 +0000 Subject: [PATCH 0277/1291] feat(ring_theory/*): generalise `minpoly.dvd` to integrally closed rings (#18021) This PR generalizes some of the theory of `minpoly` to the setting of integrally closed rings. The main results proven here are: - `minpoly.is_integrally_closed_eq_field_fractions` and variants of it - `minpoly.is_integrally_closed_dvd` - `monic.dvd_of_fraction_map_dvd_fraction_map` In a following PR, I will remove instances of `gcd_domain_dvd` (and the other results this PR generalises) from other files, and replace the file `minpoly.gcd_domain` by `minpoly.is_integrally_closed` Co-authored-by: Junyan Xu Co-authored-by: Paul Lezeau --- src/data/polynomial/algebra_map.lean | 16 ++ src/data/polynomial/monic.lean | 3 + src/data/polynomial/ring_division.lean | 28 +++ src/field_theory/minpoly/field.lean | 10 ++ src/field_theory/minpoly/gcd_monoid.lean | 178 ++++++++++++++++---- src/ring_theory/integral_closure.lean | 12 ++ src/ring_theory/polynomial/gauss_lemma.lean | 16 ++ 7 files changed, 231 insertions(+), 32 deletions(-) diff --git a/src/data/polynomial/algebra_map.lean b/src/data/polynomial/algebra_map.lean index 858805215b536..5c196b62d056b 100644 --- a/src/data/polynomial/algebra_map.lean +++ b/src/data/polynomial/algebra_map.lean @@ -251,6 +251,22 @@ lemma coeff_zero_eq_aeval_zero' (p : R[X]) : algebra_map R A (p.coeff 0) = aeval (0 : A) p := by simp [aeval_def] +lemma map_aeval_eq_aeval_map {S T U : Type*} [comm_semiring S] [comm_semiring T] [semiring U] + [algebra R S] [algebra T U] {φ : R →+* T} {ψ : S →+* U} + (h : (algebra_map T U).comp φ = ψ.comp (algebra_map R S)) (p : R[X]) (a : S) : + ψ (aeval a p) = aeval (ψ a) (p.map φ) := +begin + conv_rhs {rw [aeval_def, ← eval_map]}, + rw [map_map, h, ← map_map, eval_map, eval₂_at_apply, aeval_def, eval_map], +end + +lemma aeval_eq_zero_of_dvd_aeval_eq_zero [comm_semiring S] [comm_semiring T] [algebra S T] + {p q : S[X]} (h₁ : p ∣ q) {a : T} (h₂ : aeval a p = 0) : aeval a q = 0 := +begin + rw [aeval_def, ← eval_map] at h₂ ⊢, + exact eval_eq_zero_of_dvd_of_eval_eq_zero (polynomial.map_dvd (algebra_map S T) h₁) h₂, +end + variable (R) theorem _root_.algebra.adjoin_singleton_eq_range_aeval (x : A) : diff --git a/src/data/polynomial/monic.lean b/src/data/polynomial/monic.lean index 77cad7c1143cf..b2cf182f7f54d 100644 --- a/src/data/polynomial/monic.lean +++ b/src/data/polynomial/monic.lean @@ -346,6 +346,9 @@ begin rw [← leading_coeff_of_injective hf, hp.leading_coeff, f.map_one] end +theorem _root_.function.injective.monic_map_iff {p : R[X]} : p.monic ↔ (p.map f).monic := +⟨monic.map _, polynomial.monic_of_injective hf⟩ + end injective end semiring diff --git a/src/data/polynomial/ring_division.lean b/src/data/polynomial/ring_division.lean index 9a44d079bd618..c5e48fc31eb63 100644 --- a/src/data/polynomial/ring_division.lean +++ b/src/data/polynomial/ring_division.lean @@ -163,6 +163,34 @@ begin exact degree_le_mul_left p h2.2, end +lemma eq_zero_of_dvd_of_degree_lt {p q : R[X]} (h₁ : p ∣ q) (h₂ : degree q < degree p) : + q = 0 := +begin + by_contradiction hc, + exact (lt_iff_not_ge _ _ ).mp h₂ (degree_le_of_dvd h₁ hc), +end + +lemma eq_zero_of_dvd_of_nat_degree_lt {p q : R[X]} (h₁ : p ∣ q) (h₂ : nat_degree q < nat_degree p) : + q = 0 := +begin + by_contradiction hc, + exact (lt_iff_not_ge _ _ ).mp h₂ (nat_degree_le_of_dvd h₁ hc), +end + +theorem not_dvd_of_degree_lt {p q : R[X]} (h0 : q ≠ 0) + (hl : q.degree < p.degree) : ¬ p ∣ q := +begin + by_contra hcontra, + exact h0 (eq_zero_of_dvd_of_degree_lt hcontra hl), +end + +theorem not_dvd_of_nat_degree_lt {p q : R[X]} (h0 : q ≠ 0) + (hl : q.nat_degree < p.nat_degree) : ¬ p ∣ q := +begin + by_contra hcontra, + exact h0 (eq_zero_of_dvd_of_nat_degree_lt hcontra hl), +end + /-- This lemma is useful for working with the `int_degree` of a rational function. -/ lemma nat_degree_sub_eq_of_prod_eq {p₁ p₂ q₁ q₂ : R[X]} (hp₁ : p₁ ≠ 0) (hq₁ : q₁ ≠ 0) (hp₂ : p₂ ≠ 0) (hq₂ : q₂ ≠ 0) (h_eq : p₁ * q₂ = p₂ * q₁) : diff --git a/src/field_theory/minpoly/field.lean b/src/field_theory/minpoly/field.lean index f32ab9138dfcd..4bfcbc603bd40 100644 --- a/src/field_theory/minpoly/field.lean +++ b/src/field_theory/minpoly/field.lean @@ -84,6 +84,16 @@ lemma dvd_map_of_is_scalar_tower (A K : Type*) {R : Type*} [comm_ring A] [field minpoly K x ∣ (minpoly A x).map (algebra_map A K) := by { refine minpoly.dvd K x _, rw [aeval_map_algebra_map, minpoly.aeval] } +lemma dvd_map_of_is_scalar_tower' (R : Type*) {S : Type*} (K L : Type*) [comm_ring R] + [comm_ring S] [field K] [comm_ring L] [algebra R S] [algebra R K] [algebra S L] [algebra K L] + [algebra R L] [is_scalar_tower R K L] [is_scalar_tower R S L] (s : S): + minpoly K (algebra_map S L s) ∣ (map (algebra_map R K) (minpoly R s)) := +begin + apply minpoly.dvd K (algebra_map S L s), + rw [← map_aeval_eq_aeval_map, minpoly.aeval, map_zero], + rw [← is_scalar_tower.algebra_map_eq, ← is_scalar_tower.algebra_map_eq] +end + /-- If `y` is a conjugate of `x` over a field `K`, then it is a conjugate over a subring `R`. -/ lemma aeval_of_is_scalar_tower (R : Type*) {K T U : Type*} [comm_ring R] [field K] [comm_ring T] [algebra R K] [algebra K T] [algebra R T] [is_scalar_tower R K T] diff --git a/src/field_theory/minpoly/gcd_monoid.lean b/src/field_theory/minpoly/gcd_monoid.lean index 81719806b4065..5f14cb650f03c 100644 --- a/src/field_theory/minpoly/gcd_monoid.lean +++ b/src/field_theory/minpoly/gcd_monoid.lean @@ -24,6 +24,14 @@ This file specializes the theory of minpoly to the case of an algebra over a GCD * `gcd_domain_unique` : The minimal polynomial of an element `x` is uniquely characterized by its defining property: if there is another monic polynomial of minimal degree that has `x` as a root, then this polynomial is equal to the minimal polynomial of `x`. + + * `is_integrally_closed_eq_field_fractions`: For integrally closed domains, the minimal polynomial + over the ring is the same as the minimal polynomial over the fraction field. + +## Todo + + * Remove all results that are now special cases (e.g. we no longer need `gcd_monoid_dvd` since we + have `is_integrally_closed_dvd`). -/ open_locale classical polynomial @@ -31,20 +39,22 @@ open polynomial set function minpoly namespace minpoly -variables {R S : Type*} [comm_ring R] [comm_ring S] [is_domain R] [is_domain S] [algebra R S] +variables {R S : Type*} [comm_ring R] [comm_ring S] [is_domain R] [algebra R S] -section gcd_domain -variables (K L : Type*) [normalized_gcd_monoid R] [field K] [algebra R K] [is_fraction_ring R K] - [field L] [algebra S L] [algebra K L] [algebra R L] [is_scalar_tower R K L] - [is_scalar_tower R S L] {s : S} (hs : is_integral R s) +section + +variables (K L : Type*) [field K] [algebra R K] [is_fraction_ring R K] [field L] [algebra R L] + [algebra S L] [algebra K L] [is_scalar_tower R K L] [is_scalar_tower R S L] + +section gcd_domain -include hs +variable [normalized_gcd_monoid R] /-- For GCD domains, the minimal polynomial over the ring is the same as the minimal polynomial over the fraction field. See `minpoly.gcd_domain_eq_field_fractions'` if `S` is already a `K`-algebra. -/ -lemma gcd_domain_eq_field_fractions : +lemma gcd_domain_eq_field_fractions [is_domain S] {s : S} (hs : is_integral R s) : minpoly K (algebra_map S L s) = (minpoly R s).map (algebra_map R K) := begin refine (eq_of_irreducible_of_monic _ _ _).symm, @@ -57,8 +67,8 @@ end /-- For GCD domains, the minimal polynomial over the ring is the same as the minimal polynomial over the fraction field. Compared to `minpoly.gcd_domain_eq_field_fractions`, this version is useful if the element is in a ring that is already a `K`-algebra. -/ -lemma gcd_domain_eq_field_fractions' [algebra K S] [is_scalar_tower R K S] : - minpoly K s = (minpoly R s).map (algebra_map R K) := +lemma gcd_domain_eq_field_fractions' [is_domain S] [algebra K S] [is_scalar_tower R K S] + {s : S} (hs : is_integral R s) : minpoly K s = (minpoly R s).map (algebra_map R K) := begin let L := fraction_ring S, rw [← gcd_domain_eq_field_fractions K L hs], @@ -66,12 +76,79 @@ begin (is_integral_of_is_scalar_tower hs) rfl end -variable [no_zero_smul_divisors R S] +end gcd_domain + +variables [is_integrally_closed R] + +/-- For integrally closed domains, the minimal polynomial over the ring is the same as the minimal +polynomial over the fraction field. See `minpoly.is_integrally_closed_eq_field_fractions'` if +`S` is already a `K`-algebra. -/ +theorem is_integrally_closed_eq_field_fractions [is_domain S] {s : S} (hs : is_integral R s) : + minpoly K (algebra_map S L s) = (minpoly R s).map (algebra_map R K) := +begin + refine (eq_of_irreducible_of_monic _ _ _).symm, + { exact (polynomial.monic.irreducible_iff_irreducible_map_fraction_map + (monic hs)).1 (irreducible hs) }, + { rw [aeval_map_algebra_map, aeval_algebra_map_apply, aeval, map_zero] }, + { exact (monic hs).map _ } +end + +/-- For integrally closed domains, the minimal polynomial over the ring is the same as the minimal +polynomial over the fraction field. Compared to `minpoly.is_integrally_closed_eq_field_fractions`, +this version is useful if the element is in a ring that is already a `K`-algebra. -/ +theorem is_integrally_closed_eq_field_fractions' [is_domain S] [algebra K S] [is_scalar_tower R K S] + {s : S} (hs : is_integral R s) : minpoly K s = (minpoly R s).map (algebra_map R K) := +begin + let L := fraction_ring S, + rw [← is_integrally_closed_eq_field_fractions K L hs], + refine minpoly.eq_of_algebra_map_eq (is_fraction_ring.injective S L) + (is_integral_of_is_scalar_tower hs) rfl +end + +/-- For GCD domains, the minimal polynomial over the ring is the same as the minimal polynomial +over the fraction field. Compared to `minpoly.is_integrally_closed_eq_field_fractions`, this +version is useful if the element is in a ring that is not a domain -/ +theorem is_integrally_closed_eq_field_fractions'' [no_zero_smul_divisors S L] {s : S} + (hs : is_integral R s) : minpoly K (algebra_map S L s) = map (algebra_map R K) (minpoly R s) := +begin + --the idea of the proof is the following: since the minpoly of `a` over `Frac(R)` divides the + --minpoly of `a` over `R`, it is itself in `R`. Hence its degree is greater or equal to that of + --the minpoly of `a` over `R`. But the minpoly of `a` over `Frac(R)` divides the minpoly of a + --over `R` in `R[X]` so we are done. + + --a few "trivial" preliminary results to set up the proof + have lem0 : minpoly K (algebra_map S L s) ∣ (map (algebra_map R K) (minpoly R s)), + { exact dvd_map_of_is_scalar_tower' R K L s }, + + have lem1 : is_integral K (algebra_map S L s), + { refine is_integral_map_of_comp_eq_of_is_integral (algebra_map R K) _ _ hs, + rw [← is_scalar_tower.algebra_map_eq, ← is_scalar_tower.algebra_map_eq] }, + + obtain ⟨g, hg⟩ := is_integrally_closed.eq_map_mul_C_of_dvd K (minpoly.monic hs) lem0, + rw [(minpoly.monic lem1).leading_coeff, C_1, mul_one] at hg, + have lem2 : polynomial.aeval s g = 0, + { have := minpoly.aeval K (algebra_map S L s), + rw [← hg, ← map_aeval_eq_aeval_map, ← map_zero (algebra_map S L)] at this, + { exact no_zero_smul_divisors.algebra_map_injective S L this }, + { rw [← is_scalar_tower.algebra_map_eq, ← is_scalar_tower.algebra_map_eq] } }, + + have lem3 : g.monic, + { simpa only [function.injective.monic_map_iff (is_fraction_ring.injective R K), hg] + using minpoly.monic lem1 }, + + rw [← hg], + refine congr_arg _ (eq.symm (polynomial.eq_of_monic_of_dvd_of_nat_degree_le lem3 + (minpoly.monic hs) _ _)), + { rwa [← map_dvd_map _ (is_fraction_ring.injective R K) lem3, hg] }, + { exact nat_degree_le_nat_degree (minpoly.min R s lem3 lem2) }, +end + +end + +variables [is_domain S] [no_zero_smul_divisors R S] -/-- For GCD domains, the minimal polynomial divides any primitive polynomial that has the integral -element as root. See also `minpoly.dvd` which relaxes the assumptions on `S` in exchange for -stronger assumptions on `R`. -/ -lemma gcd_domain_dvd {P : R[X]} (hP : P ≠ 0) (hroot : polynomial.aeval s P = 0) : minpoly R s ∣ P := +lemma gcd_domain_dvd [normalized_gcd_monoid R] {s : S} (hs : is_integral R s) {P : R[X]} + (hP : P ≠ 0) (hroot : polynomial.aeval s P = 0) : minpoly R s ∣ P := begin let K := fraction_ring R, let L := fraction_ring S, @@ -87,31 +164,74 @@ begin rw [aeval_map_algebra_map, aeval_algebra_map_apply, aeval_prim_part_eq_zero hP hroot, map_zero] end +variable [is_integrally_closed R] + +/-- For integrally closed rings, the minimal polynomial divides any polynomial that has the + integral element as root. See also `minpoly.dvd` which relaxes the assumptions on `S` + in exchange for stronger assumptions on `R`. -/ +theorem is_integrally_closed_dvd [nontrivial R] (p : R[X]) {s : S} (hs : is_integral R s) : + polynomial.aeval s p = 0 ↔ minpoly R s ∣ p := +begin + refine ⟨λ hp, _, λ hp, _⟩, + + { let K := fraction_ring R, + let L := fraction_ring S, + have : minpoly K (algebra_map S L s) ∣ map (algebra_map R K) (p %ₘ (minpoly R s)), + { rw [map_mod_by_monic _ (minpoly.monic hs), mod_by_monic_eq_sub_mul_div], + refine dvd_sub (minpoly.dvd K (algebra_map S L s) _) _, + rw [← map_aeval_eq_aeval_map, hp, map_zero], + rw [← is_scalar_tower.algebra_map_eq, ← is_scalar_tower.algebra_map_eq], + + apply dvd_mul_of_dvd_left, + rw is_integrally_closed_eq_field_fractions'' K L hs, + + exact monic.map _ (minpoly.monic hs) }, + rw [is_integrally_closed_eq_field_fractions'' _ _ hs, map_dvd_map (algebra_map R K) + (is_fraction_ring.injective R K) (minpoly.monic hs)] at this, + rw [← dvd_iff_mod_by_monic_eq_zero (minpoly.monic hs)], + refine polynomial.eq_zero_of_dvd_of_degree_lt this + (degree_mod_by_monic_lt p $ minpoly.monic hs), + all_goals { apply_instance } }, + + { simpa only [ring_hom.mem_ker, ring_hom.coe_comp, coe_eval_ring_hom, + coe_map_ring_hom, function.comp_app, eval_map, ← aeval_def] using + aeval_eq_zero_of_dvd_aeval_eq_zero hp (minpoly.aeval R s) } +end + +lemma ker_eval {s : S} (hs : is_integral R s) : + ((polynomial.aeval s).to_ring_hom : R[X] →+* S).ker = ideal.span ({minpoly R s} : set R[X] ):= +begin + apply le_antisymm; intros p hp, + { rwa [ring_hom.mem_ker, alg_hom.to_ring_hom_eq_coe, alg_hom.coe_to_ring_hom, + is_integrally_closed_dvd p hs, ← ideal.mem_span_singleton] at hp }, + { rwa [ring_hom.mem_ker, alg_hom.to_ring_hom_eq_coe, alg_hom.coe_to_ring_hom, + is_integrally_closed_dvd p hs, ← ideal.mem_span_singleton] } +end + /-- If an element `x` is a root of a nonzero polynomial `p`, then the degree of `p` is at least the degree of the minimal polynomial of `x`. See also `minpoly.degree_le_of_ne_zero` which relaxes the assumptions on `S` in exchange for stronger assumptions on `R`. -/ -lemma gcd_domain_degree_le_of_ne_zero {p : R[X]} (hp0 : p ≠ 0) (hp : polynomial.aeval s p = 0) : - degree (minpoly R s) ≤ degree p := +lemma is_integrally_closed.degree_le_of_ne_zero {s : S} (hs : is_integral R s) {p : R[X]} + (hp0 : p ≠ 0) (hp : polynomial.aeval s p = 0) : degree (minpoly R s) ≤ degree p := begin rw [degree_eq_nat_degree (minpoly.ne_zero hs), degree_eq_nat_degree hp0], norm_cast, - exact nat_degree_le_of_dvd (gcd_domain_dvd hs hp0 hp) hp0 + exact nat_degree_le_of_dvd ((is_integrally_closed_dvd _ hs).mp hp) hp0 end -omit hs - /-- The minimal polynomial of an element `x` is uniquely characterized by its defining property: if there is another monic polynomial of minimal degree that has `x` as a root, then this polynomial is equal to the minimal polynomial of `x`. See also `minpoly.unique` which relaxes the assumptions on `S` in exchange for stronger assumptions on `R`. -/ -lemma gcd_domain_unique {P : R[X]} (hmo : P.monic) (hP : polynomial.aeval s P = 0) +lemma is_integrally_closed.minpoly.unique {s : S} {P : R[X]} (hmo : P.monic) + (hP : polynomial.aeval s P = 0) (Pmin : ∀ Q : R[X], Q.monic → polynomial.aeval s Q = 0 → degree P ≤ degree Q) : P = minpoly R s := begin have hs : is_integral R s := ⟨P, hmo, hP⟩, symmetry, apply eq_of_sub_eq_zero, by_contra hnz, - have := gcd_domain_degree_le_of_ne_zero hs hnz (by simp [hP]), + have := is_integrally_closed.degree_le_of_ne_zero hs hnz (by simp [hP]), contrapose! this, refine degree_sub_lt _ (ne_zero hs) _, { exact le_antisymm (min R s hmo hP) @@ -119,15 +239,13 @@ begin { rw [(monic hs).leading_coeff, hmo.leading_coeff] } end -end gcd_domain - section adjoin_root noncomputable theory open algebra polynomial adjoin_root -variables {R} {x : S} [normalized_gcd_monoid R] [no_zero_smul_divisors R S] +variables {R} {x : S} lemma to_adjoin.injective (hx : is_integral R x) : function.injective (minpoly.to_adjoin R x) := @@ -136,14 +254,10 @@ begin obtain ⟨P, hP⟩ := mk_surjective (minpoly.monic hx) P₁, by_cases hPzero : P = 0, { simpa [hPzero] using hP.symm }, - have hPcont : P.content ≠ 0 := λ h, hPzero (content_eq_zero_iff.1 h), - rw [← hP, minpoly.to_adjoin_apply', lift_hom_mk, ← subalgebra.coe_eq_zero, - aeval_subalgebra_coe, set_like.coe_mk, P.eq_C_content_mul_prim_part, aeval_mul, aeval_C] at hP₁, - replace hP₁ := eq_zero_of_ne_zero_of_mul_left_eq_zero - ((map_ne_zero_iff _ (no_zero_smul_divisors.algebra_map_injective R S)).2 hPcont) hP₁, - obtain ⟨Q, hQ⟩ := minpoly.gcd_domain_dvd hx P.is_primitive_prim_part.ne_zero hP₁, - rw [P.eq_C_content_mul_prim_part] at hP, - simpa [hQ] using hP.symm + rw [← hP, minpoly.to_adjoin_apply', lift_hom_mk, ← subalgebra.coe_eq_zero, + aeval_subalgebra_coe, set_like.coe_mk, is_integrally_closed_dvd _ hx] at hP₁, + obtain ⟨Q, hQ⟩ := hP₁, + rw [← hP, hQ, ring_hom.map_mul, mk_self, zero_mul], end /-- The algebra isomorphism `adjoin_root (minpoly R x) ≃ₐ[R] adjoin R x` -/ diff --git a/src/ring_theory/integral_closure.lean b/src/ring_theory/integral_closure.lean index ffc6cb0bbd557..49b9575bfd544 100644 --- a/src/ring_theory/integral_closure.lean +++ b/src/ring_theory/integral_closure.lean @@ -125,6 +125,18 @@ begin aeval_alg_hom_apply, aeval_map_algebra_map, aeval_def, hP.2, _root_.map_zero] end +lemma is_integral_map_of_comp_eq_of_is_integral {R S T U : Type*} [comm_ring R] [comm_ring S] + [comm_ring T] [comm_ring U] [algebra R S] [algebra T U] (φ : R →+* T) (ψ : S →+* U) + (h : (algebra_map T U).comp φ = ψ.comp (algebra_map R S)) {a : S} (ha : is_integral R a) : + is_integral T (ψ a) := +begin + rw [is_integral, ring_hom.is_integral_elem] at ⊢ ha, + obtain ⟨p, hp⟩ := ha, + refine ⟨p.map φ, hp.left.map _, _⟩, + rw [← eval_map, map_map, h, ← map_map, eval_map, eval₂_at_apply, + eval_map, hp.right, ring_hom.map_zero], +end + theorem is_integral_alg_hom_iff {A B : Type*} [ring A] [ring B] [algebra R A] [algebra R B] (f : A →ₐ[R] B) (hf : function.injective f) {x : A} : is_integral R (f x) ↔ is_integral R x := begin diff --git a/src/ring_theory/polynomial/gauss_lemma.lean b/src/ring_theory/polynomial/gauss_lemma.lean index 50fbc6b2bc3fb..6df229789277c 100644 --- a/src/ring_theory/polynomial/gauss_lemma.lean +++ b/src/ring_theory/polynomial/gauss_lemma.lean @@ -129,6 +129,22 @@ begin rw [is_root, eval_map, ← aeval_def, minpoly.aeval R x] }, end +theorem monic.dvd_of_fraction_map_dvd_fraction_map [is_integrally_closed R] {p q : R[X]} + (hp : p.monic ) (hq : q.monic) (h : q.map (algebra_map R K) ∣ p.map (algebra_map R K)) : q ∣ p := +begin + obtain ⟨r, hr⟩ := h, + obtain ⟨d', hr'⟩ := is_integrally_closed.eq_map_mul_C_of_dvd K hp (dvd_of_mul_left_eq _ hr.symm), + rw [monic.leading_coeff, C_1, mul_one] at hr', + rw [← hr', ← polynomial.map_mul] at hr, + exact dvd_of_mul_right_eq _ (polynomial.map_injective _ (is_fraction_ring.injective R K) hr.symm), + { exact monic.of_mul_monic_left (hq.map (algebra_map R K)) (by simpa [←hr] using hp.map _) }, +end + +theorem monic.dvd_iff_fraction_map_dvd_fraction_map [is_integrally_closed R] {p q : R[X]} + (hp : p.monic ) (hq : q.monic) : q.map (algebra_map R K) ∣ p.map (algebra_map R K) ↔ q ∣ p := +⟨λ h, hp.dvd_of_fraction_map_dvd_fraction_map hq h, + λ ⟨a,b⟩, ⟨a.map (algebra_map R K), b.symm ▸ polynomial.map_mul (algebra_map R K)⟩⟩ + end is_integrally_closed open is_localization From 6890b009baeb97cc338358340f80fbdfb8673055 Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Wed, 18 Jan 2023 07:04:59 +0000 Subject: [PATCH 0278/1291] feat(archive/imo/imo2019_q2): IMO 2019 Q2 (#17993) Add an IMO geometry problem to the mathlib archive. The module doc string discusses the conventions followed to translate the informal statement to a formal one (which should probably end up somewhere more generally discussing conventions for formal statements of olympiad problems, but it seems reasonable to have it here for now). This formalization uses a structure describing a configuration satisfying the conditions of the problem, to build up information gradually through many lemmas without needing to pass lots of hypotheses around (and to handle explicitly the symmetry in the configuration of the problem), which seems likely to be a convenient way to handle many olympiad geometry solutions. --- archive/imo/imo2019_q2.lean | 591 ++++++++++++++++++++++++++++++++++++ 1 file changed, 591 insertions(+) create mode 100644 archive/imo/imo2019_q2.lean diff --git a/archive/imo/imo2019_q2.lean b/archive/imo/imo2019_q2.lean new file mode 100644 index 0000000000000..de839c4451984 --- /dev/null +++ b/archive/imo/imo2019_q2.lean @@ -0,0 +1,591 @@ +/- +Copyright (c) 2022 Joseph Myers. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joseph Myers +-/ +import geometry.euclidean.angle.sphere + +/-! +# IMO 2019 Q2 + +In triangle `ABC`, point `A₁` lies on side `BC` and point `B₁` lies on side `AC`. Let `P` and +`Q` be points on segments `AA₁` and `BB₁`, respectively, such that `PQ` is parallel to `AB`. +Let `P₁` be a point on line `PB₁`, such that `B₁` lies strictly between `P` and `P₁`, and +`∠PP₁C = ∠BAC`. Similarly, let `Q₁` be a point on line `QA₁`, such that `A₁` lies strictly +between `Q` and `Q₁`, and `∠CQ₁Q = ∠CBA`. + +Prove that points `P`, `Q`, `P₁`, and `Q₁` are concyclic. + +We follow Solution 1 from the +[official solutions](https://www.imo2019.uk/wp-content/uploads/2018/07/solutions-r856.pdf). +Letting the rays `AA₁` and `BB₁` intersect the circumcircle of `ABC` at `A₂` and `B₂` +respectively, we show with an angle chase that `P`, `Q`, `A₂`, `B₂` are concyclic and let `ω` be +the circle through those points. We then show that `C`, `Q₁`, `A₂`, `A₁` are concyclic, and +then that `Q₁` lies on `ω`, and similarly that `P₁` lies on `ω`, so the required four points are +concyclic. + +Note that most of the formal proof is actually proving nondegeneracy conditions needed for that +angle chase / concyclicity argument, where an informal solution doesn't discuss those conditions +at all. Also note that (as described in `geometry.euclidean.angle.oriented.basic`) the oriented +angles used are modulo `2 * π`, so parts of the angle chase that are only valid for angles modulo +`π` (as used in the informal solution) are represented as equalities of twice angles, which we write +as `(2 : ℤ) • ∡ _ _ _ = (2 : ℤ) • _ _ _`. +-/ + +/-- +We apply the following conventions for formalizing IMO geometry problems. A problem is assumed +to take place in the plane unless that is clearly not intended, so it is not required to prove +that the points are coplanar (whether or not that in fact follows from the other conditions). +Angles in problem statements are taken to be unoriented. A reference to an angle `∠XYZ` is taken +to imply that `X` and `Z` are not equal to `Y`, since choices of junk values play no role in +informal mathematics, and those implications are included as hypotheses for the problem whether +or not they follow from the other hypotheses. Similar, a reference to `XY` as a line is taken to +imply that `X` does not equal `Y` and that is included as a hypothesis, and a reference to `XY` +being parallel to something is considered a reference to it as a line. However, such an implicit +hypothesis about two points being different is included only once for any given two points (even +if it follows from more than one reference to a line or an angle), if `X ≠ Y` is included then +`Y ≠ X` is not included separately, and such hypotheses are not included in the case where there +is also a reference in the problem to a triangle including those two points, or to strict +betweenness of three points including those two. If betweenness is stated, it is taken to be +strict betweenness. However, segments and sides are taken to include their endpoints (unless +this makes a problem false), although those degenerate cases might not necessarily have been +considered when the problem was formulated and contestants might not have been expected to deal +with them. A reference to a point being on a side or a segment is expressed directly with `wbtw` +rather than more literally with `affine_segment`. +-/ +library_note "IMO geometry formalization conventions" + +noncomputable theory + +open affine affine.simplex euclidean_geometry finite_dimensional +open_locale affine euclidean_geometry real + +local attribute [instance] fact_finite_dimensional_of_finrank_eq_succ + +variables (V : Type*) (Pt : Type*) [inner_product_space ℝ V] [metric_space Pt] +variables [normed_add_torsor V Pt] [hd2 : fact (finrank ℝ V = 2)] +include hd2 + +/-- A configuration satisfying the conditions of the problem. We define this structure to avoid +passing many hypotheses around as we build up information about the configuration; the final +result for a statement of the problem not using this structure is then deduced from one in terms +of this structure. -/ +@[nolint has_nonempty_instance] +structure imo2019q2_cfg := +(A B C A₁ B₁ P Q P₁ Q₁ : Pt) +(affine_independent_ABC : affine_independent ℝ ![A, B, C]) +(wbtw_B_A₁_C : wbtw ℝ B A₁ C) +(wbtw_A_B₁_C : wbtw ℝ A B₁ C) +(wbtw_A_P_A₁ : wbtw ℝ A P A₁) +(wbtw_B_Q_B₁ : wbtw ℝ B Q B₁) +(PQ_parallel_AB : line[ℝ, P, Q] ∥ line[ℝ, A, B]) +-- A hypothesis implicit in the named line. +(P_ne_Q : P ≠ Q) +(sbtw_P_B₁_P₁ : sbtw ℝ P B₁ P₁) +(angle_PP₁C_eq_angle_BAC : ∠ P P₁ C = ∠ B A C) +-- A hypothesis implicit in the first named angle. +(C_ne_P₁ : C ≠ P₁) +(sbtw_Q_A₁_Q₁ : sbtw ℝ Q A₁ Q₁) +(angle_CQ₁Q_eq_angle_CBA : ∠ C Q₁ Q = ∠ C B A) +-- A hypothesis implicit in the first named angle. +(C_ne_Q₁ : C ≠ Q₁) + +/-- A default choice of orientation, for lemmas that need to pick one. -/ +def some_orientation : module.oriented ℝ V (fin 2) := +⟨basis.orientation (fin_basis_of_finrank_eq _ _ hd2.out)⟩ + +variables {V Pt} + +namespace imo2019q2_cfg + +variables (cfg : imo2019q2_cfg V Pt) + +/-- The configuration has symmetry, allowing results proved for one point to be applied for +another (where the informal solution says "similarly"). -/ +def symm : imo2019q2_cfg V Pt := +{ A := cfg.B, + B := cfg.A, + C := cfg.C, + A₁ := cfg.B₁, + B₁ := cfg.A₁, + P := cfg.Q, + Q := cfg.P, + P₁ := cfg.Q₁, + Q₁ := cfg.P₁, + affine_independent_ABC := begin + rw ←affine_independent_equiv (equiv.swap (0 : fin 3) 1), + convert cfg.affine_independent_ABC using 1, + ext x, + fin_cases x; + refl + end, + wbtw_B_A₁_C := cfg.wbtw_A_B₁_C, + wbtw_A_B₁_C := cfg.wbtw_B_A₁_C, + wbtw_A_P_A₁ := cfg.wbtw_B_Q_B₁, + wbtw_B_Q_B₁ := cfg.wbtw_A_P_A₁, + PQ_parallel_AB := set.pair_comm cfg.P cfg.Q ▸ set.pair_comm cfg.A cfg.B ▸ cfg.PQ_parallel_AB, + P_ne_Q := cfg.P_ne_Q.symm, + sbtw_P_B₁_P₁ := cfg.sbtw_Q_A₁_Q₁, + angle_PP₁C_eq_angle_BAC := + angle_comm cfg.C cfg.Q₁ cfg.Q ▸ angle_comm cfg.C cfg.B cfg.A ▸ cfg.angle_CQ₁Q_eq_angle_CBA, + C_ne_P₁ := cfg.C_ne_Q₁, + sbtw_Q_A₁_Q₁ := cfg.sbtw_P_B₁_P₁, + angle_CQ₁Q_eq_angle_CBA := + angle_comm cfg.P cfg.P₁ cfg.C ▸ angle_comm cfg.B cfg.A cfg.C ▸ cfg.angle_PP₁C_eq_angle_BAC, + C_ne_Q₁ := cfg.C_ne_P₁ } + +/-! ### Configuration properties that are obvious from the diagram, and construction of the +points `A₂` and `B₂` -/ + +lemma A_ne_B : cfg.A ≠ cfg.B := cfg.affine_independent_ABC.injective.ne + (dec_trivial : (0 : fin 3) ≠ 1) + +lemma A_ne_C : cfg.A ≠ cfg.C := cfg.affine_independent_ABC.injective.ne + (dec_trivial : (0 : fin 3) ≠ 2) + +lemma B_ne_C : cfg.B ≠ cfg.C := cfg.affine_independent_ABC.injective.ne + (dec_trivial : (1 : fin 3) ≠ 2) + +lemma not_collinear_ABC : ¬collinear ℝ ({cfg.A, cfg.B, cfg.C} : set Pt) := +affine_independent_iff_not_collinear_set.1 cfg.affine_independent_ABC + +/-- `ABC` as a `triangle`. -/ +def triangle_ABC : triangle ℝ Pt := ⟨_, cfg.affine_independent_ABC⟩ + +lemma A_mem_circumsphere : cfg.A ∈ cfg.triangle_ABC.circumsphere := +cfg.triangle_ABC.mem_circumsphere 0 + +lemma B_mem_circumsphere : cfg.B ∈ cfg.triangle_ABC.circumsphere := +cfg.triangle_ABC.mem_circumsphere 1 + +lemma C_mem_circumsphere : cfg.C ∈ cfg.triangle_ABC.circumsphere := +cfg.triangle_ABC.mem_circumsphere 2 + +lemma symm_triangle_ABC : cfg.symm.triangle_ABC = cfg.triangle_ABC.reindex (equiv.swap 0 1) := +by { ext i, fin_cases i; refl } + +lemma symm_triangle_ABC_circumsphere : + cfg.symm.triangle_ABC.circumsphere = cfg.triangle_ABC.circumsphere := +by rw [symm_triangle_ABC, affine.simplex.circumsphere_reindex] + +/-- `A₂` is the second point of intersection of the ray `AA₁` with the circumcircle of `ABC`. -/ +def A₂ : Pt := cfg.triangle_ABC.circumsphere.second_inter cfg.A (cfg.A₁ -ᵥ cfg.A) + +/-- `B₂` is the second point of intersection of the ray `BB₁` with the circumcircle of `ABC`. -/ +def B₂ : Pt := cfg.triangle_ABC.circumsphere.second_inter cfg.B (cfg.B₁ -ᵥ cfg.B) + +lemma A₂_mem_circumsphere : cfg.A₂ ∈ cfg.triangle_ABC.circumsphere := +(sphere.second_inter_mem _).2 cfg.A_mem_circumsphere + +lemma B₂_mem_circumsphere : cfg.B₂ ∈ cfg.triangle_ABC.circumsphere := +(sphere.second_inter_mem _).2 cfg.B_mem_circumsphere + +lemma symm_A₂ : cfg.symm.A₂ = cfg.B₂ := +by { simp_rw [A₂, B₂, symm_triangle_ABC_circumsphere], refl } + +lemma QP_parallel_BA : line[ℝ, cfg.Q, cfg.P] ∥ line[ℝ, cfg.B, cfg.A] := +by { rw [set.pair_comm cfg.Q, set.pair_comm cfg.B], exact cfg.PQ_parallel_AB } + +lemma A_ne_A₁ : cfg.A ≠ cfg.A₁ := +begin + intro h, + have h' := cfg.not_collinear_ABC, + rw [h, set.insert_comm] at h', + exact h' cfg.wbtw_B_A₁_C.collinear +end + +lemma collinear_PAA₁A₂ : collinear ℝ ({cfg.P, cfg.A, cfg.A₁, cfg.A₂} : set Pt) := +begin + rw [A₂, + (cfg.triangle_ABC.circumsphere.second_inter_collinear cfg.A cfg.A₁).collinear_insert_iff_of_ne + (set.mem_insert _ _) (set.mem_insert_of_mem _ (set.mem_insert _ _)) cfg.A_ne_A₁, + set.insert_comm], + exact cfg.wbtw_A_P_A₁.collinear +end + +lemma A₁_ne_C : cfg.A₁ ≠ cfg.C := +begin + intro h, + have hsbtw := cfg.sbtw_Q_A₁_Q₁, + rw h at hsbtw, + have ha := hsbtw.angle₂₃₁_eq_zero, + rw [angle_CQ₁Q_eq_angle_CBA, angle_comm] at ha, + exact (angle_ne_zero_of_not_collinear cfg.not_collinear_ABC) ha +end + +lemma B₁_ne_C : cfg.B₁ ≠ cfg.C := cfg.symm.A₁_ne_C + +lemma Q_not_mem_CB : cfg.Q ∉ line[ℝ, cfg.C, cfg.B] := +begin + intro hQ, + have hQA₁ : line[ℝ, cfg.Q, cfg.A₁] ≤ line[ℝ, cfg.C, cfg.B] := + affine_span_pair_le_of_mem_of_mem hQ cfg.wbtw_B_A₁_C.symm.mem_affine_span, + have hQ₁ : cfg.Q₁ ∈ line[ℝ, cfg.C, cfg.B], + { rw affine_subspace.le_def' at hQA₁, + exact hQA₁ _ cfg.sbtw_Q_A₁_Q₁.right_mem_affine_span }, + have hc : collinear ℝ ({cfg.C, cfg.Q₁, cfg.Q} : set Pt), + { have hc' : collinear ℝ ({cfg.B, cfg.C, cfg.Q₁, cfg.Q} : set Pt), + { rw [set.insert_comm cfg.B, set.insert_comm cfg.B, set.pair_comm, set.insert_comm cfg.C, + set.insert_comm cfg.C], + exact collinear_insert_insert_of_mem_affine_span_pair hQ₁ hQ }, + exact hc'.subset (set.subset_insert _ _) }, + rw [collinear_iff_eq_or_eq_or_angle_eq_zero_or_angle_eq_pi, cfg.angle_CQ₁Q_eq_angle_CBA, + or_iff_right cfg.C_ne_Q₁, or_iff_right cfg.sbtw_Q_A₁_Q₁.left_ne_right, angle_comm] at hc, + exact cfg.not_collinear_ABC (hc.elim collinear_of_angle_eq_zero collinear_of_angle_eq_pi) +end + +lemma Q_ne_B : cfg.Q ≠ cfg.B := +begin + intro h, + have h' := cfg.Q_not_mem_CB, + rw h at h', + exact h' (right_mem_affine_span_pair _ _ _) +end + +lemma s_opp_side_CB_Q_Q₁ : line[ℝ, cfg.C, cfg.B].s_opp_side cfg.Q cfg.Q₁ := +cfg.sbtw_Q_A₁_Q₁.s_opp_side_of_not_mem_of_mem cfg.Q_not_mem_CB cfg.wbtw_B_A₁_C.symm.mem_affine_span + +/-! ### Relate the orientations of different angles in the configuration -/ + +section oriented + +variables [module.oriented ℝ V (fin 2)] + +lemma oangle_CQ₁Q_sign_eq_oangle_CBA_sign : + (∡ cfg.C cfg.Q₁ cfg.Q).sign = (∡ cfg.C cfg.B cfg.A).sign := +by rw [←cfg.sbtw_Q_A₁_Q₁.symm.oangle_eq_right, + cfg.s_opp_side_CB_Q_Q₁.oangle_sign_eq_neg (left_mem_affine_span_pair ℝ cfg.C cfg.B) + cfg.wbtw_B_A₁_C.symm.mem_affine_span, ←real.angle.sign_neg, ←oangle_rev, + cfg.wbtw_B_A₁_C.oangle_sign_eq_of_ne_right cfg.Q cfg.A₁_ne_C, oangle_rotate_sign, + cfg.wbtw_B_Q_B₁.oangle_eq_right cfg.Q_ne_B, + cfg.wbtw_A_B₁_C.symm.oangle_sign_eq_of_ne_left cfg.B cfg.B₁_ne_C.symm] + +lemma oangle_CQ₁Q_eq_oangle_CBA : ∡ cfg.C cfg.Q₁ cfg.Q = ∡ cfg.C cfg.B cfg.A := +oangle_eq_of_angle_eq_of_sign_eq cfg.angle_CQ₁Q_eq_angle_CBA cfg.oangle_CQ₁Q_sign_eq_oangle_CBA_sign + +end oriented + +/-! ### More obvious configuration properties -/ + +lemma A₁_ne_B : cfg.A₁ ≠ cfg.B := +begin + intro h, + have hwbtw := cfg.wbtw_A_P_A₁, + rw h at hwbtw, + have hPQ : line[ℝ, cfg.P, cfg.Q] = line[ℝ, cfg.A, cfg.B], + { rw affine_subspace.eq_iff_direction_eq_of_mem (left_mem_affine_span_pair _ _ _) + hwbtw.mem_affine_span, + exact cfg.PQ_parallel_AB.direction_eq }, + haveI := some_orientation V, + have haQ : (2 : ℤ) • ∡ cfg.C cfg.B cfg.Q = (2 : ℤ) • ∡ cfg.C cfg.B cfg.A, + { rw [collinear.two_zsmul_oangle_eq_right _ cfg.A_ne_B cfg.Q_ne_B], + rw [set.pair_comm, set.insert_comm], + refine collinear_insert_of_mem_affine_span_pair _, + rw ←hPQ, + exact right_mem_affine_span_pair _ _ _ }, + have ha : (2 : ℤ) • ∡ cfg.C cfg.B cfg.Q = (2 : ℤ) • ∡ cfg.C cfg.Q₁ cfg.Q, + { rw [oangle_CQ₁Q_eq_oangle_CBA, haQ] }, + have hn : ¬collinear ℝ ({cfg.C, cfg.B, cfg.Q} : set Pt), + { rw [collinear_iff_of_two_zsmul_oangle_eq haQ, set.pair_comm, set.insert_comm, set.pair_comm], + exact cfg.not_collinear_ABC }, + have hc := cospherical_of_two_zsmul_oangle_eq_of_not_collinear ha hn, + have hBQ₁ : cfg.B ≠ cfg.Q₁, { rw [←h], exact cfg.sbtw_Q_A₁_Q₁.ne_right }, + have hQQ₁ : cfg.Q ≠ cfg.Q₁ := cfg.sbtw_Q_A₁_Q₁.left_ne_right, + have hBQ₁Q : affine_independent ℝ ![cfg.B, cfg.Q₁, cfg.Q] := + hc.affine_independent_of_mem_of_ne (set.mem_insert_of_mem _ (set.mem_insert _ _)) + (set.mem_insert_of_mem _ (set.mem_insert_of_mem _ (set.mem_insert _ _))) + (set.mem_insert_of_mem _ (set.mem_insert_of_mem _ (set.mem_insert_of_mem _ + (set.mem_singleton _)))) hBQ₁ cfg.Q_ne_B.symm hQQ₁.symm, + rw affine_independent_iff_not_collinear_set at hBQ₁Q, + refine hBQ₁Q _, + rw [←h, set.pair_comm, set.insert_comm], + exact cfg.sbtw_Q_A₁_Q₁.wbtw.collinear +end + +lemma sbtw_B_A₁_C : sbtw ℝ cfg.B cfg.A₁ cfg.C := ⟨cfg.wbtw_B_A₁_C, cfg.A₁_ne_B, cfg.A₁_ne_C⟩ + +lemma sbtw_A_B₁_C : sbtw ℝ cfg.A cfg.B₁ cfg.C := cfg.symm.sbtw_B_A₁_C + +lemma sbtw_A_A₁_A₂ : sbtw ℝ cfg.A cfg.A₁ cfg.A₂ := +begin + refine sphere.sbtw_second_inter cfg.A_mem_circumsphere _, + convert cfg.sbtw_B_A₁_C.dist_lt_max_dist _, + change _ = max (dist (cfg.triangle_ABC.points 1) _) (dist (cfg.triangle_ABC.points 2) _), + simp_rw [circumsphere_center, circumsphere_radius, dist_circumcenter_eq_circumradius, max_self] +end + +lemma sbtw_B_B₁_B₂ : sbtw ℝ cfg.B cfg.B₁ cfg.B₂ := +by { rw ←cfg.symm_A₂, exact cfg.symm.sbtw_A_A₁_A₂ } + +lemma A₂_ne_A : cfg.A₂ ≠ cfg.A := cfg.sbtw_A_A₁_A₂.left_ne_right.symm + +lemma A₂_ne_P : cfg.A₂ ≠ cfg.P := (cfg.sbtw_A_A₁_A₂.trans_wbtw_left_ne cfg.wbtw_A_P_A₁).symm + +lemma A₂_ne_B : cfg.A₂ ≠ cfg.B := +begin + intro h, + have h₁ := cfg.sbtw_A_A₁_A₂, + rw h at h₁, + refine cfg.not_collinear_ABC _, + have hc : collinear ℝ ({cfg.A, cfg.C, cfg.B, cfg.A₁} : set Pt) := + collinear_insert_insert_of_mem_affine_span_pair h₁.left_mem_affine_span + cfg.sbtw_B_A₁_C.right_mem_affine_span, + refine hc.subset _, + rw [set.pair_comm _ cfg.A₁, set.insert_comm _ cfg.A₁, set.insert_comm _ cfg.A₁, set.pair_comm], + exact set.subset_insert _ _ +end + +lemma A₂_ne_C : cfg.A₂ ≠ cfg.C := +begin + intro h, + have h₁ := cfg.sbtw_A_A₁_A₂, + rw h at h₁, + refine cfg.not_collinear_ABC _, + have hc : collinear ℝ ({cfg.A, cfg.B, cfg.C, cfg.A₁} : set Pt) := + collinear_insert_insert_of_mem_affine_span_pair h₁.left_mem_affine_span + cfg.sbtw_B_A₁_C.left_mem_affine_span, + refine hc.subset (set.insert_subset_insert (set.insert_subset_insert _)), + rw set.singleton_subset_iff, + exact set.mem_insert _ _ +end + +lemma B₂_ne_B : cfg.B₂ ≠ cfg.B := by { rw ←symm_A₂, exact cfg.symm.A₂_ne_A } + +lemma B₂_ne_Q : cfg.B₂ ≠ cfg.Q := by { rw ←symm_A₂, exact cfg.symm.A₂_ne_P } + +lemma B₂_ne_A₂ : cfg.B₂ ≠ cfg.A₂ := +begin + intro h, + have hA : sbtw ℝ (cfg.triangle_ABC.points 1) cfg.A₁ (cfg.triangle_ABC.points 2) := + cfg.sbtw_B_A₁_C, + have hB : sbtw ℝ (cfg.triangle_ABC.points 0) cfg.B₁ (cfg.triangle_ABC.points 2) := + cfg.sbtw_A_B₁_C, + have hA' : cfg.A₂ ∈ line[ℝ, cfg.triangle_ABC.points 0, cfg.A₁] := + sphere.second_inter_vsub_mem_affine_span _ _ _, + have hB' : cfg.A₂ ∈ line[ℝ, cfg.triangle_ABC.points 1, cfg.B₁], + { rw ←h, exact sphere.second_inter_vsub_mem_affine_span _ _ _ }, + exact (sbtw_of_sbtw_of_sbtw_of_mem_affine_span_pair dec_trivial hA hB hA' hB').symm.not_rotate + cfg.sbtw_A_A₁_A₂.wbtw +end + +lemma wbtw_B_Q_B₂ : wbtw ℝ cfg.B cfg.Q cfg.B₂ := cfg.sbtw_B_B₁_B₂.wbtw.trans_left cfg.wbtw_B_Q_B₁ + +/-! ### The first equality in the first angle chase in the solution -/ + +section oriented + +variables [module.oriented ℝ V (fin 2)] + +lemma two_zsmul_oangle_QPA₂_eq_two_zsmul_oangle_BAA₂ : + (2 : ℤ) • ∡ cfg.Q cfg.P cfg.A₂ = (2 : ℤ) • ∡ cfg.B cfg.A cfg.A₂ := +begin + refine two_zsmul_oangle_of_parallel cfg.QP_parallel_BA _, + convert affine_subspace.parallel.refl _ using 1, + rw [cfg.collinear_PAA₁A₂.affine_span_eq_of_ne + (set.mem_insert_of_mem _ (set.mem_insert_of_mem _ (set.mem_insert_of_mem _ + (set.mem_singleton _)))) + (set.mem_insert_of_mem _ (set.mem_insert _ _)) cfg.A₂_ne_A, + cfg.collinear_PAA₁A₂.affine_span_eq_of_ne + (set.mem_insert_of_mem _ (set.mem_insert_of_mem _ (set.mem_insert_of_mem _ + (set.mem_singleton _)))) + (set.mem_insert _ _) cfg.A₂_ne_P] +end + +end oriented + +/-! ### More obvious configuration properties -/ + +lemma not_collinear_QPA₂ : ¬ collinear ℝ ({cfg.Q, cfg.P, cfg.A₂} : set Pt) := +begin + haveI := some_orientation V, + rw [collinear_iff_of_two_zsmul_oangle_eq cfg.two_zsmul_oangle_QPA₂_eq_two_zsmul_oangle_BAA₂, + ←affine_independent_iff_not_collinear_set], + have h : cospherical ({cfg.B, cfg.A, cfg.A₂} : set Pt), + { refine cfg.triangle_ABC.circumsphere.cospherical.subset _, + simp [set.insert_subset, cfg.A_mem_circumsphere, cfg.B_mem_circumsphere, + cfg.A₂_mem_circumsphere] }, + exact h.affine_independent_of_ne cfg.A_ne_B.symm cfg.A₂_ne_B.symm cfg.A₂_ne_A.symm +end + +lemma Q₁_ne_A₂ : cfg.Q₁ ≠ cfg.A₂ := +begin + intro h, + have h₁ := cfg.sbtw_Q_A₁_Q₁, + rw h at h₁, + refine cfg.not_collinear_QPA₂ _, + have hA₂ := cfg.sbtw_A_A₁_A₂.right_mem_affine_span, + have hA₂A₁ : line[ℝ, cfg.A₂, cfg.A₁] ≤ line[ℝ, cfg.A, cfg.A₁] := + affine_span_pair_le_of_left_mem hA₂, + have hQ : cfg.Q ∈ line[ℝ, cfg.A, cfg.A₁], + { rw affine_subspace.le_def' at hA₂A₁, + exact hA₂A₁ _ h₁.left_mem_affine_span }, + exact collinear_triple_of_mem_affine_span_pair hQ cfg.wbtw_A_P_A₁.mem_affine_span hA₂ +end + +lemma affine_independent_QPA₂ : affine_independent ℝ ![cfg.Q, cfg.P, cfg.A₂] := +affine_independent_iff_not_collinear_set.2 cfg.not_collinear_QPA₂ + +lemma affine_independent_PQB₂ : affine_independent ℝ ![cfg.P, cfg.Q, cfg.B₂] := +by { rw ←symm_A₂, exact cfg.symm.affine_independent_QPA₂ } + +/-- `QPA₂` as a `triangle`. -/ +def triangle_QPA₂ : triangle ℝ Pt := ⟨_, cfg.affine_independent_QPA₂⟩ + +/-- `PQB₂` as a `triangle`. -/ +def triangle_PQB₂ : triangle ℝ Pt := ⟨_, cfg.affine_independent_PQB₂⟩ + +lemma symm_triangle_QPA₂ : cfg.symm.triangle_QPA₂ = cfg.triangle_PQB₂ := +by { simp_rw [triangle_PQB₂, ←symm_A₂], ext i, fin_cases i; refl } + +/-- `ω` is the circle containing `Q`, `P` and `A₂`, which will be shown also to contain `B₂`, +`P₁` and `Q₁`. -/ +def ω : sphere Pt := cfg.triangle_QPA₂.circumsphere + +lemma P_mem_ω : cfg.P ∈ cfg.ω := cfg.triangle_QPA₂.mem_circumsphere 1 + +lemma Q_mem_ω : cfg.Q ∈ cfg.ω := cfg.triangle_QPA₂.mem_circumsphere 0 + +/-! ### The rest of the first angle chase in the solution -/ + +section oriented + +variables [module.oriented ℝ V (fin 2)] + +lemma two_zsmul_oangle_QPA₂_eq_two_zsmul_oangle_QB₂A₂ : + (2 : ℤ) • ∡ cfg.Q cfg.P cfg.A₂ = (2 : ℤ) • ∡ cfg.Q cfg.B₂ cfg.A₂ := +calc (2 : ℤ) • ∡ cfg.Q cfg.P cfg.A₂ = (2 : ℤ) • ∡ cfg.B cfg.A cfg.A₂ : + cfg.two_zsmul_oangle_QPA₂_eq_two_zsmul_oangle_BAA₂ + ... = (2 : ℤ) • ∡ cfg.B cfg.B₂ cfg.A₂ : + sphere.two_zsmul_oangle_eq cfg.B_mem_circumsphere cfg.A_mem_circumsphere + cfg.B₂_mem_circumsphere cfg.A₂_mem_circumsphere cfg.A_ne_B cfg.A₂_ne_A.symm + cfg.B₂_ne_B cfg.B₂_ne_A₂ + ... = (2 : ℤ) • ∡ cfg.Q cfg.B₂ cfg.A₂ : + by rw cfg.wbtw_B_Q_B₂.symm.oangle_eq_left cfg.B₂_ne_Q.symm + +end oriented + +/-! ### Conclusions from that first angle chase -/ + +lemma cospherical_QPB₂A₂ : cospherical ({cfg.Q, cfg.P, cfg.B₂, cfg.A₂} : set Pt) := +begin + haveI := some_orientation V, + exact cospherical_of_two_zsmul_oangle_eq_of_not_collinear + cfg.two_zsmul_oangle_QPA₂_eq_two_zsmul_oangle_QB₂A₂ cfg.not_collinear_QPA₂ +end + +lemma symm_ω_eq_triangle_PQB₂_circumsphere : cfg.symm.ω = cfg.triangle_PQB₂.circumsphere := +by rw [ω, symm_triangle_QPA₂] + +lemma symm_ω : cfg.symm.ω = cfg.ω := +begin + rw [symm_ω_eq_triangle_PQB₂_circumsphere, ω], + refine circumsphere_eq_of_cospherical hd2.out cfg.cospherical_QPB₂A₂ _ _, + { simp only [triangle_PQB₂, matrix.range_cons, matrix.range_empty, set.singleton_union, + insert_emptyc_eq], + rw set.insert_comm, + refine set.insert_subset_insert (set.insert_subset_insert _), + simp }, + { simp only [triangle_QPA₂, matrix.range_cons, matrix.range_empty, set.singleton_union, + insert_emptyc_eq], + refine set.insert_subset_insert (set.insert_subset_insert _), + simp } +end + +/-! ### The second angle chase in the solution -/ + +section oriented + +variables [module.oriented ℝ V (fin 2)] + +lemma two_zsmul_oangle_CA₂A₁_eq_two_zsmul_oangle_CBA : + (2 : ℤ) • ∡ cfg.C cfg.A₂ cfg.A₁ = (2 : ℤ) • ∡ cfg.C cfg.B cfg.A := +calc (2 : ℤ) • ∡ cfg.C cfg.A₂ cfg.A₁ = (2 : ℤ) • ∡ cfg.C cfg.A₂ cfg.A : + by rw cfg.sbtw_A_A₁_A₂.symm.oangle_eq_right + ... = (2 : ℤ) • ∡ cfg.C cfg.B cfg.A : + sphere.two_zsmul_oangle_eq cfg.C_mem_circumsphere cfg.A₂_mem_circumsphere + cfg.B_mem_circumsphere cfg.A_mem_circumsphere cfg.A₂_ne_C cfg.A₂_ne_A cfg.B_ne_C + cfg.A_ne_B.symm + +lemma two_zsmul_oangle_CA₂A₁_eq_two_zsmul_oangle_CQ₁A₁ : + (2 : ℤ) • ∡ cfg.C cfg.A₂ cfg.A₁ = (2 : ℤ) • ∡ cfg.C cfg.Q₁ cfg.A₁ := +calc (2 : ℤ) • ∡ cfg.C cfg.A₂ cfg.A₁ = (2 : ℤ) • ∡ cfg.C cfg.B cfg.A : + cfg.two_zsmul_oangle_CA₂A₁_eq_two_zsmul_oangle_CBA + ... = (2 : ℤ) • ∡ cfg.C cfg.Q₁ cfg.Q : by rw oangle_CQ₁Q_eq_oangle_CBA + ... = (2 : ℤ) • ∡ cfg.C cfg.Q₁ cfg.A₁ : by rw cfg.sbtw_Q_A₁_Q₁.symm.oangle_eq_right + +end oriented + +/-! ### Conclusions from that second angle chase -/ + +lemma not_collinear_CA₂A₁ : ¬collinear ℝ ({cfg.C, cfg.A₂, cfg.A₁} : set Pt) := +begin + haveI := some_orientation V, + rw [collinear_iff_of_two_zsmul_oangle_eq cfg.two_zsmul_oangle_CA₂A₁_eq_two_zsmul_oangle_CBA, + set.pair_comm, set.insert_comm, set.pair_comm], + exact cfg.not_collinear_ABC +end + +lemma cospherical_A₁Q₁CA₂ : cospherical ({cfg.A₁, cfg.Q₁, cfg.C, cfg.A₂} : set Pt) := +begin + haveI := some_orientation V, + rw [set.insert_comm cfg.Q₁, set.insert_comm cfg.A₁, set.pair_comm, set.insert_comm cfg.A₁, + set.pair_comm], + exact cospherical_of_two_zsmul_oangle_eq_of_not_collinear + cfg.two_zsmul_oangle_CA₂A₁_eq_two_zsmul_oangle_CQ₁A₁ cfg.not_collinear_CA₂A₁ +end + +/-! ### The third angle chase in the solution -/ + +section oriented + +variables [module.oriented ℝ V (fin 2)] + +lemma two_zsmul_oangle_QQ₁A₂_eq_two_zsmul_oangle_QPA₂ : + (2 : ℤ) • ∡ cfg.Q cfg.Q₁ cfg.A₂ = (2 : ℤ) • ∡ cfg.Q cfg.P cfg.A₂ := +calc (2 : ℤ) • ∡ cfg.Q cfg.Q₁ cfg.A₂ = (2 : ℤ) • ∡ cfg.A₁ cfg.Q₁ cfg.A₂ : + by rw cfg.sbtw_Q_A₁_Q₁.symm.oangle_eq_left + ... = (2 : ℤ) • ∡ cfg.A₁ cfg.C cfg.A₂ : + cfg.cospherical_A₁Q₁CA₂.two_zsmul_oangle_eq cfg.sbtw_Q_A₁_Q₁.right_ne cfg.Q₁_ne_A₂ + cfg.A₁_ne_C.symm cfg.A₂_ne_C.symm + ... = (2 : ℤ) • ∡ cfg.B cfg.C cfg.A₂ : by rw cfg.sbtw_B_A₁_C.symm.oangle_eq_left + ... = (2 : ℤ) • ∡ cfg.B cfg.A cfg.A₂ : + sphere.two_zsmul_oangle_eq cfg.B_mem_circumsphere cfg.C_mem_circumsphere + cfg.A_mem_circumsphere cfg.A₂_mem_circumsphere cfg.B_ne_C.symm cfg.A₂_ne_C.symm cfg.A_ne_B + cfg.A₂_ne_A.symm + ... = (2 : ℤ) • ∡ cfg.Q cfg.P cfg.A₂ : cfg.two_zsmul_oangle_QPA₂_eq_two_zsmul_oangle_BAA₂.symm + +end oriented + +/-! ### Conclusions from that third angle chase -/ + +lemma Q₁_mem_ω : cfg.Q₁ ∈ cfg.ω := +begin + haveI := some_orientation V, + exact affine.triangle.mem_circumsphere_of_two_zsmul_oangle_eq (dec_trivial : (0 : fin 3) ≠ 1) + (dec_trivial : (0 : fin 3) ≠ 2) dec_trivial cfg.two_zsmul_oangle_QQ₁A₂_eq_two_zsmul_oangle_QPA₂ +end + +lemma P₁_mem_ω : cfg.P₁ ∈ cfg.ω := by { rw ←symm_ω, exact cfg.symm.Q₁_mem_ω } + +theorem result : concyclic ({cfg.P, cfg.Q, cfg.P₁, cfg.Q₁} : set Pt) := +begin + refine ⟨_, coplanar_of_fact_finrank_eq_two _⟩, + rw cospherical_iff_exists_sphere, + refine ⟨cfg.ω, _⟩, + simp only [set.insert_subset, set.singleton_subset_iff], + exact ⟨cfg.P_mem_ω, cfg.Q_mem_ω, cfg.P₁_mem_ω, cfg.Q₁_mem_ω⟩ +end + +end imo2019q2_cfg + +theorem imo2019_q2 (A B C A₁ B₁ P Q P₁ Q₁ : Pt) + (affine_independent_ABC : affine_independent ℝ ![A, B, C]) + (wbtw_B_A₁_C : wbtw ℝ B A₁ C) (wbtw_A_B₁_C : wbtw ℝ A B₁ C) (wbtw_A_P_A₁ : wbtw ℝ A P A₁) + (wbtw_B_Q_B₁ : wbtw ℝ B Q B₁) (PQ_parallel_AB : line[ℝ, P, Q] ∥ line[ℝ, A, B]) (P_ne_Q : P ≠ Q) + (sbtw_P_B₁_P₁ : sbtw ℝ P B₁ P₁) (angle_PP₁C_eq_angle_BAC : ∠ P P₁ C = ∠ B A C) + (C_ne_P₁ : C ≠ P₁) (sbtw_Q_A₁_Q₁ : sbtw ℝ Q A₁ Q₁) + (angle_CQ₁Q_eq_angle_CBA : ∠ C Q₁ Q = ∠ C B A) (C_ne_Q₁ : C ≠ Q₁) : + concyclic ({P, Q, P₁, Q₁} : set Pt) := +(⟨A, B, C, A₁, B₁, P, Q, P₁, Q₁, affine_independent_ABC, wbtw_B_A₁_C, wbtw_A_B₁_C, wbtw_A_P_A₁, + wbtw_B_Q_B₁, PQ_parallel_AB, P_ne_Q, sbtw_P_B₁_P₁, angle_PP₁C_eq_angle_BAC, C_ne_P₁, + sbtw_Q_A₁_Q₁, angle_CQ₁Q_eq_angle_CBA, C_ne_Q₁⟩ : imo2019q2_cfg V Pt).result From 26703ef268fe458f540b3bcd5adcb38d923ba9e7 Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Wed, 18 Jan 2023 09:22:55 +0000 Subject: [PATCH 0279/1291] refactor(analysis/normed_space/operator_norm): split (#18205) The file `analysis/normed_space/operator_norm` is now nearly 2000 lines long, and it's a pretty heavy import since it now contains everything on the strong topologies, etc. This PR splits out 250 lines: - some lemmas about infinite sums under continuous linear maps to a new file `topology/algebra/module/infinite_sum` (they were already stated for topological vector spaces, so definitely in the wrong location before) - some constructors for continuous linear maps between normed spaces to a new file `analysis/normed_space/continuous_linear_map`, which is intended to be a more lightweight import Later I hope to refactor some other files (e.g. `is_R_or_C`, `analysis/complex/basic`) to import only the new file `analysis/normed_space/continuous_linear_map`, not `analysis/normed_space/operator_norm`. --- src/analysis/analytic/basic.lean | 1 + .../normed_space/continuous_linear_map.lean | 269 ++++++++++++++++++ .../normed_space/finite_dimension.lean | 1 + src/analysis/normed_space/operator_norm.lean | 244 +--------------- src/topology/algebra/module/infinite_sum.lean | 70 +++++ 5 files changed, 345 insertions(+), 240 deletions(-) create mode 100644 src/analysis/normed_space/continuous_linear_map.lean create mode 100644 src/topology/algebra/module/infinite_sum.lean diff --git a/src/analysis/analytic/basic.lean b/src/analysis/analytic/basic.lean index d8f3bfea0cfeb..3dd9a5016a666 100644 --- a/src/analysis/analytic/basic.lean +++ b/src/analysis/analytic/basic.lean @@ -6,6 +6,7 @@ Authors: Sébastien Gouëzel, Yury Kudryashov import analysis.calculus.formal_multilinear_series import analysis.specific_limits.normed import logic.equiv.fin +import topology.algebra.module.infinite_sum /-! # Analytic functions diff --git a/src/analysis/normed_space/continuous_linear_map.lean b/src/analysis/normed_space/continuous_linear_map.lean new file mode 100644 index 0000000000000..bb3bc4ae41afb --- /dev/null +++ b/src/analysis/normed_space/continuous_linear_map.lean @@ -0,0 +1,269 @@ +/- +Copyright (c) 2019 Jan-David Salchow. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jan-David Salchow, Sébastien Gouëzel, Jean Lo +-/ +import analysis.normed_space.basic + +/-! # Constructions of continuous linear maps between (semi-)normed spaces + +A fundamental fact about (semi-)linear maps between normed spaces over sensible fields is that +continuity and boundedness are equivalent conditions. That is, for normed spaces `E`, `F`, a +`linear_map` `f : E →ₛₗ[σ] F` is the coercion of some `continuous_linear_map` `f' : E →SL[σ] F`, if +and only if there exists a bound `C` such that for all `x`, `‖f x‖ ≤ C * ‖x‖`. + +We prove one direction in this file: `linear_map.mk_continuous`, boundedness implies continuity. The +other direction, `continuous_linear_map.bound`, is deferred to a later file, where the +strong operator topology on `E →SL[σ] F` is available, because it is natural to use +`continuous_linear_map.bound` to define a norm `⨆ x, ‖f x‖ / ‖x‖` on `E →SL[σ] F` and to show that +this is compatible with the strong operator topology. + +This file also contains several corollaries of `linear_map.mk_continuous`: other "easy" +constructions of continuous linear maps between normed spaces. + +This file is meant to be lightweight (it is imported by much of the analysis library); think twice +before adding imports! +-/ + +open metric continuous_linear_map +open set real + +open_locale nnreal + +variables {𝕜 𝕜₂ E F G : Type*} + +variables [normed_field 𝕜] [normed_field 𝕜₂] + +/-! # General constructions -/ + +section seminormed + +variables [seminormed_add_comm_group E] [seminormed_add_comm_group F] [seminormed_add_comm_group G] +variables [normed_space 𝕜 E] [normed_space 𝕜₂ F] [normed_space 𝕜 G] +variables {σ : 𝕜 →+* 𝕜₂} (f : E →ₛₗ[σ] F) + +/-- Construct a continuous linear map from a linear map and a bound on this linear map. +The fact that the norm of the continuous linear map is then controlled is given in +`linear_map.mk_continuous_norm_le`. -/ +def linear_map.mk_continuous (C : ℝ) (h : ∀x, ‖f x‖ ≤ C * ‖x‖) : E →SL[σ] F := +⟨f, add_monoid_hom_class.continuous_of_bound f C h⟩ + +/-- Reinterpret a linear map `𝕜 →ₗ[𝕜] E` as a continuous linear map. This construction +is generalized to the case of any finite dimensional domain +in `linear_map.to_continuous_linear_map`. -/ +def linear_map.to_continuous_linear_map₁ (f : 𝕜 →ₗ[𝕜] E) : 𝕜 →L[𝕜] E := +f.mk_continuous (‖f 1‖) $ λ x, le_of_eq $ +by { conv_lhs { rw ← mul_one x }, rw [← smul_eq_mul, f.map_smul, norm_smul, mul_comm] } + +/-- Construct a continuous linear map from a linear map and the existence of a bound on this linear +map. If you have an explicit bound, use `linear_map.mk_continuous` instead, as a norm estimate will +follow automatically in `linear_map.mk_continuous_norm_le`. -/ +def linear_map.mk_continuous_of_exists_bound (h : ∃C, ∀x, ‖f x‖ ≤ C * ‖x‖) : E →SL[σ] F := +⟨f, let ⟨C, hC⟩ := h in add_monoid_hom_class.continuous_of_bound f C hC⟩ + +lemma continuous_of_linear_of_boundₛₗ {f : E → F} (h_add : ∀ x y, f (x + y) = f x + f y) + (h_smul : ∀ (c : 𝕜) x, f (c • x) = (σ c) • f x) {C : ℝ} (h_bound : ∀ x, ‖f x‖ ≤ C*‖x‖) : + continuous f := +let φ : E →ₛₗ[σ] F := { to_fun := f, map_add' := h_add, map_smul' := h_smul } in +add_monoid_hom_class.continuous_of_bound φ C h_bound + +lemma continuous_of_linear_of_bound {f : E → G} (h_add : ∀ x y, f (x + y) = f x + f y) + (h_smul : ∀ (c : 𝕜) x, f (c • x) = c • f x) {C : ℝ} (h_bound : ∀ x, ‖f x‖ ≤ C*‖x‖) : + continuous f := +let φ : E →ₗ[𝕜] G := { to_fun := f, map_add' := h_add, map_smul' := h_smul } in +add_monoid_hom_class.continuous_of_bound φ C h_bound + +@[simp, norm_cast] lemma linear_map.mk_continuous_coe (C : ℝ) (h : ∀x, ‖f x‖ ≤ C * ‖x‖) : + ((f.mk_continuous C h) : E →ₛₗ[σ] F) = f := rfl + +@[simp] lemma linear_map.mk_continuous_apply (C : ℝ) (h : ∀x, ‖f x‖ ≤ C * ‖x‖) (x : E) : + f.mk_continuous C h x = f x := rfl + +@[simp, norm_cast] lemma linear_map.mk_continuous_of_exists_bound_coe + (h : ∃C, ∀x, ‖f x‖ ≤ C * ‖x‖) : + ((f.mk_continuous_of_exists_bound h) : E →ₛₗ[σ] F) = f := rfl + +@[simp] lemma linear_map.mk_continuous_of_exists_bound_apply (h : ∃C, ∀x, ‖f x‖ ≤ C * ‖x‖) (x : E) : + f.mk_continuous_of_exists_bound h x = f x := rfl + +@[simp] lemma linear_map.to_continuous_linear_map₁_coe (f : 𝕜 →ₗ[𝕜] E) : + (f.to_continuous_linear_map₁ : 𝕜 →ₗ[𝕜] E) = f := +rfl + +@[simp] lemma linear_map.to_continuous_linear_map₁_apply (f : 𝕜 →ₗ[𝕜] E) (x) : + f.to_continuous_linear_map₁ x = f x := +rfl + +namespace continuous_linear_map + +theorem antilipschitz_of_bound (f : E →SL[σ] F) {K : ℝ≥0} (h : ∀ x, ‖x‖ ≤ K * ‖f x‖) : + antilipschitz_with K f := +add_monoid_hom_class.antilipschitz_of_bound _ h + +lemma bound_of_antilipschitz (f : E →SL[σ] F) {K : ℝ≥0} (h : antilipschitz_with K f) (x) : + ‖x‖ ≤ K * ‖f x‖ := +add_monoid_hom_class.bound_of_antilipschitz _ h x + +end continuous_linear_map + +section + +variables {σ₂₁ : 𝕜₂ →+* 𝕜} [ring_hom_inv_pair σ σ₂₁] [ring_hom_inv_pair σ₂₁ σ] + +include σ₂₁ + +/-- Construct a continuous linear equivalence from a linear equivalence together with +bounds in both directions. -/ +def linear_equiv.to_continuous_linear_equiv_of_bounds (e : E ≃ₛₗ[σ] F) (C_to C_inv : ℝ) + (h_to : ∀ x, ‖e x‖ ≤ C_to * ‖x‖) (h_inv : ∀ x : F, ‖e.symm x‖ ≤ C_inv * ‖x‖) : E ≃SL[σ] F := +{ to_linear_equiv := e, + continuous_to_fun := add_monoid_hom_class.continuous_of_bound e C_to h_to, + continuous_inv_fun := add_monoid_hom_class.continuous_of_bound e.symm C_inv h_inv } + +end + +end seminormed + +section normed + +variables [normed_add_comm_group E] [normed_add_comm_group F] [normed_space 𝕜 E] [normed_space 𝕜₂ F] +variables {σ : 𝕜 →+* 𝕜₂} (f g : E →SL[σ] F) (x y z : E) + +theorem continuous_linear_map.uniform_embedding_of_bound {K : ℝ≥0} (hf : ∀ x, ‖x‖ ≤ K * ‖f x‖) : + uniform_embedding f := +(add_monoid_hom_class.antilipschitz_of_bound f hf).uniform_embedding f.uniform_continuous + +end normed + +/-! ## Homotheties -/ + +section seminormed + +variables [seminormed_add_comm_group E] [seminormed_add_comm_group F] +variables [normed_space 𝕜 E] [normed_space 𝕜₂ F] +variables {σ : 𝕜 →+* 𝕜₂} (f : E →ₛₗ[σ] F) + +/-- A (semi-)linear map which is a homothety is a continuous linear map. + Since the field `𝕜` need not have `ℝ` as a subfield, this theorem is not directly deducible from + the corresponding theorem about isometries plus a theorem about scalar multiplication. Likewise + for the other theorems about homotheties in this file. + -/ +def continuous_linear_map.of_homothety (f : E →ₛₗ[σ] F) (a : ℝ) (hf : ∀x, ‖f x‖ = a * ‖x‖) : + E →SL[σ] F := +f.mk_continuous a (λ x, le_of_eq (hf x)) + +variables {σ₂₁ : 𝕜₂ →+* 𝕜} [ring_hom_inv_pair σ σ₂₁] [ring_hom_inv_pair σ₂₁ σ] + +include σ₂₁ + +lemma continuous_linear_equiv.homothety_inverse (a : ℝ) (ha : 0 < a) (f : E ≃ₛₗ[σ] F) : + (∀ (x : E), ‖f x‖ = a * ‖x‖) → (∀ (y : F), ‖f.symm y‖ = a⁻¹ * ‖y‖) := +begin + intros hf y, + calc ‖(f.symm) y‖ = a⁻¹ * (a * ‖ (f.symm) y‖) : _ + ... = a⁻¹ * ‖f ((f.symm) y)‖ : by rw hf + ... = a⁻¹ * ‖y‖ : by simp, + rw [← mul_assoc, inv_mul_cancel (ne_of_lt ha).symm, one_mul], +end + +/-- A linear equivalence which is a homothety is a continuous linear equivalence. -/ +noncomputable def continuous_linear_equiv.of_homothety (f : E ≃ₛₗ[σ] F) (a : ℝ) (ha : 0 < a) + (hf : ∀x, ‖f x‖ = a * ‖x‖) : + E ≃SL[σ] F := +linear_equiv.to_continuous_linear_equiv_of_bounds f a a⁻¹ + (λ x, (hf x).le) (λ x, (continuous_linear_equiv.homothety_inverse a ha f hf x).le) + +end seminormed + +/- ## The span of a single vector -/ + +section seminormed + +variables [seminormed_add_comm_group E] [normed_space 𝕜 E] + +namespace continuous_linear_map + +variable (𝕜) + +lemma to_span_singleton_homothety (x : E) (c : 𝕜) : + ‖linear_map.to_span_singleton 𝕜 E x c‖ = ‖x‖ * ‖c‖ := +by {rw mul_comm, exact norm_smul _ _} + +/-- Given an element `x` of a normed space `E` over a field `𝕜`, the natural continuous + linear map from `𝕜` to `E` by taking multiples of `x`.-/ +def to_span_singleton (x : E) : 𝕜 →L[𝕜] E := +of_homothety (linear_map.to_span_singleton 𝕜 E x) ‖x‖ (to_span_singleton_homothety 𝕜 x) + +lemma to_span_singleton_apply (x : E) (r : 𝕜) : to_span_singleton 𝕜 x r = r • x := +by simp [to_span_singleton, of_homothety, linear_map.to_span_singleton] + +lemma to_span_singleton_add (x y : E) : + to_span_singleton 𝕜 (x + y) = to_span_singleton 𝕜 x + to_span_singleton 𝕜 y := +by { ext1, simp [to_span_singleton_apply], } + +lemma to_span_singleton_smul' (𝕜') [normed_field 𝕜'] [normed_space 𝕜' E] + [smul_comm_class 𝕜 𝕜' E] (c : 𝕜') (x : E) : + to_span_singleton 𝕜 (c • x) = c • to_span_singleton 𝕜 x := +by { ext1, rw [to_span_singleton_apply, smul_apply, to_span_singleton_apply, smul_comm], } + +lemma to_span_singleton_smul (c : 𝕜) (x : E) : + to_span_singleton 𝕜 (c • x) = c • to_span_singleton 𝕜 x := +to_span_singleton_smul' 𝕜 𝕜 c x + +end continuous_linear_map + +section + +namespace continuous_linear_equiv + +variable (𝕜) + +lemma to_span_nonzero_singleton_homothety (x : E) (h : x ≠ 0) (c : 𝕜) : + ‖linear_equiv.to_span_nonzero_singleton 𝕜 E x h c‖ = ‖x‖ * ‖c‖ := +continuous_linear_map.to_span_singleton_homothety _ _ _ + +end continuous_linear_equiv + +end + +end seminormed + +section normed + +variables [normed_add_comm_group E] [normed_space 𝕜 E] + +namespace continuous_linear_equiv +variable (𝕜) + +/-- Given a nonzero element `x` of a normed space `E₁` over a field `𝕜`, the natural + continuous linear equivalence from `E₁` to the span of `x`.-/ +noncomputable def to_span_nonzero_singleton (x : E) (h : x ≠ 0) : 𝕜 ≃L[𝕜] (𝕜 ∙ x) := +of_homothety + (linear_equiv.to_span_nonzero_singleton 𝕜 E x h) + ‖x‖ + (norm_pos_iff.mpr h) + (to_span_nonzero_singleton_homothety 𝕜 x h) + +/-- Given a nonzero element `x` of a normed space `E₁` over a field `𝕜`, the natural continuous + linear map from the span of `x` to `𝕜`.-/ +noncomputable def coord (x : E) (h : x ≠ 0) : (𝕜 ∙ x) →L[𝕜] 𝕜 := + (to_span_nonzero_singleton 𝕜 x h).symm + +@[simp] lemma coe_to_span_nonzero_singleton_symm {x : E} (h : x ≠ 0) : + ⇑(to_span_nonzero_singleton 𝕜 x h).symm = coord 𝕜 x h := rfl + +@[simp] lemma coord_to_span_nonzero_singleton {x : E} (h : x ≠ 0) (c : 𝕜) : + coord 𝕜 x h (to_span_nonzero_singleton 𝕜 x h c) = c := +(to_span_nonzero_singleton 𝕜 x h).symm_apply_apply c + +@[simp] lemma to_span_nonzero_singleton_coord {x : E} (h : x ≠ 0) (y : 𝕜 ∙ x) : + to_span_nonzero_singleton 𝕜 x h (coord 𝕜 x h y) = y := +(to_span_nonzero_singleton 𝕜 x h).apply_symm_apply y + +@[simp] lemma coord_self (x : E) (h : x ≠ 0) : + (coord 𝕜 x h) (⟨x, submodule.mem_span_singleton_self x⟩ : 𝕜 ∙ x) = 1 := +linear_equiv.coord_self 𝕜 E x h + +end continuous_linear_equiv + +end normed diff --git a/src/analysis/normed_space/finite_dimension.lean b/src/analysis/normed_space/finite_dimension.lean index 87caaba6ba994..79ddedfebaee7 100644 --- a/src/analysis/normed_space/finite_dimension.lean +++ b/src/analysis/normed_space/finite_dimension.lean @@ -10,6 +10,7 @@ import analysis.normed_space.operator_norm import analysis.normed_space.riesz_lemma import linear_algebra.matrix.to_lin import topology.algebra.module.finite_dimension +import topology.algebra.module.infinite_sum import topology.instances.matrix /-! diff --git a/src/analysis/normed_space/operator_norm.lean b/src/analysis/normed_space/operator_norm.lean index 364d681e6c05b..915303ee6f361 100644 --- a/src/analysis/normed_space/operator_norm.lean +++ b/src/analysis/normed_space/operator_norm.lean @@ -5,6 +5,7 @@ Authors: Jan-David Salchow, Sébastien Gouëzel, Jean Lo -/ import algebra.algebra.tower import analysis.asymptotics.asymptotics +import analysis.normed_space.continuous_linear_map import analysis.normed_space.linear_isometry import topology.algebra.module.strong_topology @@ -30,74 +31,10 @@ variables {𝕜 𝕜₂ 𝕜₃ E Eₗ F Fₗ G Gₗ 𝓕 : Type*} section semi_normed -variables [seminormed_add_comm_group E] [seminormed_add_comm_group Eₗ] [seminormed_add_comm_group F] - [seminormed_add_comm_group Fₗ] [seminormed_add_comm_group G] [seminormed_add_comm_group Gₗ] - open metric continuous_linear_map -section normed_field -/-! Most statements in this file require the field to be non-discrete, -as this is necessary to deduce an inequality `‖f x‖ ≤ C ‖x‖` from the continuity of f. -However, the other direction always holds. -In this section, we just assume that `𝕜` is a normed field. -In the remainder of the file, it will be non-discrete. -/ - -variables [normed_field 𝕜] [normed_field 𝕜₂] [normed_space 𝕜 E] [normed_space 𝕜₂ F] -variables [normed_space 𝕜 G] {σ : 𝕜 →+* 𝕜₂} (f : E →ₛₗ[σ] F) - -/-- Construct a continuous linear map from a linear map and a bound on this linear map. -The fact that the norm of the continuous linear map is then controlled is given in -`linear_map.mk_continuous_norm_le`. -/ -def linear_map.mk_continuous (C : ℝ) (h : ∀x, ‖f x‖ ≤ C * ‖x‖) : E →SL[σ] F := -⟨f, add_monoid_hom_class.continuous_of_bound f C h⟩ - -/-- Reinterpret a linear map `𝕜 →ₗ[𝕜] E` as a continuous linear map. This construction -is generalized to the case of any finite dimensional domain -in `linear_map.to_continuous_linear_map`. -/ -def linear_map.to_continuous_linear_map₁ (f : 𝕜 →ₗ[𝕜] E) : 𝕜 →L[𝕜] E := -f.mk_continuous (‖f 1‖) $ λ x, le_of_eq $ -by { conv_lhs { rw ← mul_one x }, rw [← smul_eq_mul, f.map_smul, norm_smul, mul_comm] } - -/-- Construct a continuous linear map from a linear map and the existence of a bound on this linear -map. If you have an explicit bound, use `linear_map.mk_continuous` instead, as a norm estimate will -follow automatically in `linear_map.mk_continuous_norm_le`. -/ -def linear_map.mk_continuous_of_exists_bound (h : ∃C, ∀x, ‖f x‖ ≤ C * ‖x‖) : E →SL[σ] F := -⟨f, let ⟨C, hC⟩ := h in add_monoid_hom_class.continuous_of_bound f C hC⟩ - -lemma continuous_of_linear_of_boundₛₗ {f : E → F} (h_add : ∀ x y, f (x + y) = f x + f y) - (h_smul : ∀ (c : 𝕜) x, f (c • x) = (σ c) • f x) {C : ℝ} (h_bound : ∀ x, ‖f x‖ ≤ C*‖x‖) : - continuous f := -let φ : E →ₛₗ[σ] F := { to_fun := f, map_add' := h_add, map_smul' := h_smul } in -add_monoid_hom_class.continuous_of_bound φ C h_bound - -lemma continuous_of_linear_of_bound {f : E → G} (h_add : ∀ x y, f (x + y) = f x + f y) - (h_smul : ∀ (c : 𝕜) x, f (c • x) = c • f x) {C : ℝ} (h_bound : ∀ x, ‖f x‖ ≤ C*‖x‖) : - continuous f := -let φ : E →ₗ[𝕜] G := { to_fun := f, map_add' := h_add, map_smul' := h_smul } in -add_monoid_hom_class.continuous_of_bound φ C h_bound - -@[simp, norm_cast] lemma linear_map.mk_continuous_coe (C : ℝ) (h : ∀x, ‖f x‖ ≤ C * ‖x‖) : - ((f.mk_continuous C h) : E →ₛₗ[σ] F) = f := rfl - -@[simp] lemma linear_map.mk_continuous_apply (C : ℝ) (h : ∀x, ‖f x‖ ≤ C * ‖x‖) (x : E) : - f.mk_continuous C h x = f x := rfl - -@[simp, norm_cast] lemma linear_map.mk_continuous_of_exists_bound_coe - (h : ∃C, ∀x, ‖f x‖ ≤ C * ‖x‖) : - ((f.mk_continuous_of_exists_bound h) : E →ₛₗ[σ] F) = f := rfl - -@[simp] lemma linear_map.mk_continuous_of_exists_bound_apply (h : ∃C, ∀x, ‖f x‖ ≤ C * ‖x‖) (x : E) : - f.mk_continuous_of_exists_bound h x = f x := rfl - -@[simp] lemma linear_map.to_continuous_linear_map₁_coe (f : 𝕜 →ₗ[𝕜] E) : - (f.to_continuous_linear_map₁ : 𝕜 →ₗ[𝕜] E) = f := -rfl - -@[simp] lemma linear_map.to_continuous_linear_map₁_apply (f : 𝕜 →ₗ[𝕜] E) (x) : - f.to_continuous_linear_map₁ x = f x := -rfl - -end normed_field +variables [seminormed_add_comm_group E] [seminormed_add_comm_group Eₗ] [seminormed_add_comm_group F] + [seminormed_add_comm_group Fₗ] [seminormed_add_comm_group G] [seminormed_add_comm_group Gₗ] variables [nontrivially_normed_field 𝕜] [nontrivially_normed_field 𝕜₂] [nontrivially_normed_field 𝕜₃] [normed_space 𝕜 E] [normed_space 𝕜 Eₗ] [normed_space 𝕜₂ F] @@ -164,41 +101,6 @@ semilinear_map_class.bound_of_continuous f f.2 section open filter -/-- A linear map which is a homothety is a continuous linear map. - Since the field `𝕜` need not have `ℝ` as a subfield, this theorem is not directly deducible from - the corresponding theorem about isometries plus a theorem about scalar multiplication. Likewise - for the other theorems about homotheties in this file. - -/ -def of_homothety (f : E →ₛₗ[σ₁₂] F) (a : ℝ) (hf : ∀x, ‖f x‖ = a * ‖x‖) : E →SL[σ₁₂] F := -f.mk_continuous a (λ x, le_of_eq (hf x)) - -variable (𝕜) - -lemma to_span_singleton_homothety (x : E) (c : 𝕜) : - ‖linear_map.to_span_singleton 𝕜 E x c‖ = ‖x‖ * ‖c‖ := -by {rw mul_comm, exact norm_smul _ _} - -/-- Given an element `x` of a normed space `E` over a field `𝕜`, the natural continuous - linear map from `𝕜` to `E` by taking multiples of `x`.-/ -def to_span_singleton (x : E) : 𝕜 →L[𝕜] E := -of_homothety (linear_map.to_span_singleton 𝕜 E x) ‖x‖ (to_span_singleton_homothety 𝕜 x) - -lemma to_span_singleton_apply (x : E) (r : 𝕜) : to_span_singleton 𝕜 x r = r • x := -by simp [to_span_singleton, of_homothety, linear_map.to_span_singleton] - -lemma to_span_singleton_add (x y : E) : - to_span_singleton 𝕜 (x + y) = to_span_singleton 𝕜 x + to_span_singleton 𝕜 y := -by { ext1, simp [to_span_singleton_apply], } - -lemma to_span_singleton_smul' (𝕜') [normed_field 𝕜'] [normed_space 𝕜' E] - [smul_comm_class 𝕜 𝕜' E] (c : 𝕜') (x : E) : - to_span_singleton 𝕜 (c • x) = c • to_span_singleton 𝕜 x := -by { ext1, rw [to_span_singleton_apply, smul_apply, to_span_singleton_apply, smul_comm], } - -lemma to_span_singleton_smul (c : 𝕜) (x : E) : - to_span_singleton 𝕜 (c • x) = c • to_span_singleton 𝕜 x := -to_span_singleton_smul' 𝕜 𝕜 c x - variables (𝕜 E) /-- Given a unit-length element `x` of a normed space `E` over a field `𝕜`, the natural linear isometry map from `𝕜` to `E` by taking multiples of `x`.-/ @@ -261,14 +163,6 @@ le_antisymm (φ.op_norm_le_bound M_nonneg h_above) lemma op_norm_neg (f : E →SL[σ₁₂] F) : ‖-f‖ = ‖f‖ := by simp only [norm_def, neg_apply, norm_neg] -theorem antilipschitz_of_bound (f : E →SL[σ₁₂] F) {K : ℝ≥0} (h : ∀ x, ‖x‖ ≤ K * ‖f x‖) : - antilipschitz_with K f := -add_monoid_hom_class.antilipschitz_of_bound _ h - -lemma bound_of_antilipschitz (f : E →SL[σ₁₂] F) {K : ℝ≥0} (h : antilipschitz_with K f) (x) : - ‖x‖ ≤ K * ‖f x‖ := -add_monoid_hom_class.bound_of_antilipschitz _ h x - section variables [ring_hom_isometric σ₁₂] [ring_hom_isometric σ₂₃] @@ -1153,67 +1047,6 @@ K.subtypeₗᵢ.norm_to_continuous_linear_map_le end submodule -section has_sum - --- Results in this section hold for continuous additive monoid homomorphisms or equivalences but we --- don't have bundled continuous additive homomorphisms. - -variables {ι R R₂ M M₂ : Type*} [semiring R] [semiring R₂] [add_comm_monoid M] [module R M] - [add_comm_monoid M₂] [module R₂ M₂] [topological_space M] [topological_space M₂] - {σ : R →+* R₂} {σ' : R₂ →+* R} [ring_hom_inv_pair σ σ'] [ring_hom_inv_pair σ' σ] - -/-- Applying a continuous linear map commutes with taking an (infinite) sum. -/ -protected lemma continuous_linear_map.has_sum {f : ι → M} (φ : M →SL[σ] M₂) {x : M} - (hf : has_sum f x) : - has_sum (λ (b:ι), φ (f b)) (φ x) := -by simpa only using hf.map φ.to_linear_map.to_add_monoid_hom φ.continuous - -alias continuous_linear_map.has_sum ← has_sum.mapL - -protected lemma continuous_linear_map.summable {f : ι → M} (φ : M →SL[σ] M₂) (hf : summable f) : - summable (λ b:ι, φ (f b)) := -(hf.has_sum.mapL φ).summable - -alias continuous_linear_map.summable ← summable.mapL - -protected lemma continuous_linear_map.map_tsum [t2_space M₂] {f : ι → M} - (φ : M →SL[σ] M₂) (hf : summable f) : φ (∑' z, f z) = ∑' z, φ (f z) := -(hf.has_sum.mapL φ).tsum_eq.symm - -include σ' -/-- Applying a continuous linear map commutes with taking an (infinite) sum. -/ -protected lemma continuous_linear_equiv.has_sum {f : ι → M} (e : M ≃SL[σ] M₂) {y : M₂} : - has_sum (λ (b:ι), e (f b)) y ↔ has_sum f (e.symm y) := -⟨λ h, by simpa only [e.symm.coe_coe, e.symm_apply_apply] using h.mapL (e.symm : M₂ →SL[σ'] M), - λ h, by simpa only [e.coe_coe, e.apply_symm_apply] using (e : M →SL[σ] M₂).has_sum h⟩ - -/-- Applying a continuous linear map commutes with taking an (infinite) sum. -/ -protected lemma continuous_linear_equiv.has_sum' {f : ι → M} (e : M ≃SL[σ] M₂) {x : M} : - has_sum (λ (b:ι), e (f b)) (e x) ↔ has_sum f x := -by rw [e.has_sum, continuous_linear_equiv.symm_apply_apply] - -protected lemma continuous_linear_equiv.summable {f : ι → M} (e : M ≃SL[σ] M₂) : - summable (λ b:ι, e (f b)) ↔ summable f := -⟨λ hf, (e.has_sum.1 hf.has_sum).summable, (e : M →SL[σ] M₂).summable⟩ - - -lemma continuous_linear_equiv.tsum_eq_iff [t2_space M] [t2_space M₂] {f : ι → M} - (e : M ≃SL[σ] M₂) {y : M₂} : ∑' z, e (f z) = y ↔ ∑' z, f z = e.symm y := -begin - by_cases hf : summable f, - { exact ⟨λ h, (e.has_sum.mp ((e.summable.mpr hf).has_sum_iff.mpr h)).tsum_eq, - λ h, (e.has_sum.mpr (hf.has_sum_iff.mpr h)).tsum_eq⟩ }, - { have hf' : ¬summable (λ z, e (f z)) := λ h, hf (e.summable.mp h), - rw [tsum_eq_zero_of_not_summable hf, tsum_eq_zero_of_not_summable hf'], - exact ⟨by { rintro rfl, simp }, λ H, by simpa using (congr_arg (λ z, e z) H)⟩ } -end - -protected lemma continuous_linear_equiv.map_tsum [t2_space M] [t2_space M₂] {f : ι → M} - (e : M ≃SL[σ] M₂) : e (∑' z, f z) = ∑' z, e (f z) := -by { refine symm (e.tsum_eq_iff.mpr _), rw e.symm_apply_apply _ } - -end has_sum - namespace continuous_linear_equiv section @@ -1235,27 +1068,9 @@ theorem is_O_sub (l : filter E) (x : E) : (λ x', e (x' - x)) =O[l] (λ x', x' - end variables {σ₂₁ : 𝕜₂ →+* 𝕜} [ring_hom_inv_pair σ₁₂ σ₂₁] [ring_hom_inv_pair σ₂₁ σ₁₂] +variables [ring_hom_isometric σ₂₁] (e : E ≃SL[σ₁₂] F) include σ₂₁ -lemma homothety_inverse (a : ℝ) (ha : 0 < a) (f : E ≃ₛₗ[σ₁₂] F) : - (∀ (x : E), ‖f x‖ = a * ‖x‖) → (∀ (y : F), ‖f.symm y‖ = a⁻¹ * ‖y‖) := -begin - intros hf y, - calc ‖(f.symm) y‖ = a⁻¹ * (a * ‖ (f.symm) y‖) : _ - ... = a⁻¹ * ‖f ((f.symm) y)‖ : by rw hf - ... = a⁻¹ * ‖y‖ : by simp, - rw [← mul_assoc, inv_mul_cancel (ne_of_lt ha).symm, one_mul], -end - -/-- A linear equivalence which is a homothety is a continuous linear equivalence. -/ -def of_homothety (f : E ≃ₛₗ[σ₁₂] F) (a : ℝ) (ha : 0 < a) (hf : ∀x, ‖f x‖ = a * ‖x‖) : - E ≃SL[σ₁₂] F := -{ to_linear_equiv := f, - continuous_to_fun := add_monoid_hom_class.continuous_of_bound f a (λ x, le_of_eq (hf x)), - continuous_inv_fun := add_monoid_hom_class.continuous_of_bound f.symm a⁻¹ - (λ x, le_of_eq (homothety_inverse a ha f hf x)) } - -variables [ring_hom_isometric σ₂₁] (e : E ≃SL[σ₁₂] F) theorem is_O_comp_rev {α : Type*} (f : α → E) (l : filter α) : f =O[l] (λ x', e (f x')) := (e.symm.is_O_comp _ l).congr_left $ λ _, e.symm_apply_apply _ @@ -1263,28 +1078,9 @@ theorem is_O_comp_rev {α : Type*} (f : α → E) (l : filter α) : f =O[l] (λ theorem is_O_sub_rev (l : filter E) (x : E) : (λ x', x' - x) =O[l] (λ x', e (x' - x)) := e.is_O_comp_rev _ _ -omit σ₂₁ - -variable (𝕜) - -lemma to_span_nonzero_singleton_homothety (x : E) (h : x ≠ 0) (c : 𝕜) : - ‖linear_equiv.to_span_nonzero_singleton 𝕜 E x h c‖ = ‖x‖ * ‖c‖ := -continuous_linear_map.to_span_singleton_homothety _ _ _ - end continuous_linear_equiv variables {σ₂₁ : 𝕜₂ →+* 𝕜} [ring_hom_inv_pair σ₁₂ σ₂₁] [ring_hom_inv_pair σ₂₁ σ₁₂] -include σ₂₁ - -/-- Construct a continuous linear equivalence from a linear equivalence together with -bounds in both directions. -/ -def linear_equiv.to_continuous_linear_equiv_of_bounds (e : E ≃ₛₗ[σ₁₂] F) (C_to C_inv : ℝ) - (h_to : ∀ x, ‖e x‖ ≤ C_to * ‖x‖) (h_inv : ∀ x : F, ‖e.symm x‖ ≤ C_inv * ‖x‖) : E ≃SL[σ₁₂] F := -{ to_linear_equiv := e, - continuous_to_fun := add_monoid_hom_class.continuous_of_bound e C_to h_to, - continuous_inv_fun := add_monoid_hom_class.continuous_of_bound e.symm C_inv h_inv } - -omit σ₂₁ namespace continuous_linear_map variables {E' F' : Type*} [seminormed_add_comm_group E'] [seminormed_add_comm_group F'] @@ -1418,10 +1214,6 @@ end variable (f) -theorem uniform_embedding_of_bound {K : ℝ≥0} (hf : ∀ x, ‖x‖ ≤ K * ‖f x‖) : - uniform_embedding f := -(add_monoid_hom_class.antilipschitz_of_bound f hf).uniform_embedding f.uniform_continuous - /-- If a continuous linear map is a uniform embedding, then it is expands the distances by a positive factor.-/ theorem antilipschitz_of_uniform_embedding (f : E →L[𝕜] Fₗ) (hf : uniform_embedding f) : @@ -1895,30 +1687,6 @@ subsingleton_or_norm_symm_pos e variable (𝕜) -/-- Given a nonzero element `x` of a normed space `E₁` over a field `𝕜`, the natural - continuous linear equivalence from `E₁` to the span of `x`.-/ -def to_span_nonzero_singleton (x : E) (h : x ≠ 0) : 𝕜 ≃L[𝕜] (𝕜 ∙ x) := -of_homothety - (linear_equiv.to_span_nonzero_singleton 𝕜 E x h) - ‖x‖ - (norm_pos_iff.mpr h) - (to_span_nonzero_singleton_homothety 𝕜 x h) - -/-- Given a nonzero element `x` of a normed space `E₁` over a field `𝕜`, the natural continuous - linear map from the span of `x` to `𝕜`.-/ -def coord (x : E) (h : x ≠ 0) : (𝕜 ∙ x) →L[𝕜] 𝕜 := (to_span_nonzero_singleton 𝕜 x h).symm - -@[simp] lemma coe_to_span_nonzero_singleton_symm {x : E} (h : x ≠ 0) : - ⇑(to_span_nonzero_singleton 𝕜 x h).symm = coord 𝕜 x h := rfl - -@[simp] lemma coord_to_span_nonzero_singleton {x : E} (h : x ≠ 0) (c : 𝕜) : - coord 𝕜 x h (to_span_nonzero_singleton 𝕜 x h c) = c := -(to_span_nonzero_singleton 𝕜 x h).symm_apply_apply c - -@[simp] lemma to_span_nonzero_singleton_coord {x : E} (h : x ≠ 0) (y : 𝕜 ∙ x) : - to_span_nonzero_singleton 𝕜 x h (coord 𝕜 x h y) = y := -(to_span_nonzero_singleton 𝕜 x h).apply_symm_apply y - @[simp] lemma coord_norm (x : E) (h : x ≠ 0) : ‖coord 𝕜 x h‖ = ‖x‖⁻¹ := begin have hx : 0 < ‖x‖ := (norm_pos_iff.mpr h), @@ -1927,10 +1695,6 @@ begin (λ y, homothety_inverse _ hx _ (to_span_nonzero_singleton_homothety 𝕜 x h) _) end -@[simp] lemma coord_self (x : E) (h : x ≠ 0) : - (coord 𝕜 x h) (⟨x, submodule.mem_span_singleton_self x⟩ : 𝕜 ∙ x) = 1 := -linear_equiv.coord_self 𝕜 E x h - variables {𝕜} {𝕜₄ : Type*} [nontrivially_normed_field 𝕜₄] variables {H : Type*} [normed_add_comm_group H] [normed_space 𝕜₄ H] [normed_space 𝕜₃ G] variables {σ₂₃ : 𝕜₂ →+* 𝕜₃} {σ₁₃ : 𝕜 →+* 𝕜₃} diff --git a/src/topology/algebra/module/infinite_sum.lean b/src/topology/algebra/module/infinite_sum.lean new file mode 100644 index 0000000000000..4360e47ab5dd1 --- /dev/null +++ b/src/topology/algebra/module/infinite_sum.lean @@ -0,0 +1,70 @@ +/- +Copyright (c) 2020 Heather Macbeth. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Heather Macbeth, Yury Kudryashov, Frédéric Dupuis +-/ +import topology.algebra.module.basic +import topology.algebra.infinite_sum + +/-! # Infinite sums in topological vector spaces -/ + +section has_sum + +-- Results in this section hold for continuous additive monoid homomorphisms or equivalences but we +-- don't have bundled continuous additive homomorphisms. + +variables {ι R R₂ M M₂ : Type*} [semiring R] [semiring R₂] [add_comm_monoid M] [module R M] + [add_comm_monoid M₂] [module R₂ M₂] [topological_space M] [topological_space M₂] + {σ : R →+* R₂} {σ' : R₂ →+* R} [ring_hom_inv_pair σ σ'] [ring_hom_inv_pair σ' σ] + +/-- Applying a continuous linear map commutes with taking an (infinite) sum. -/ +protected lemma continuous_linear_map.has_sum {f : ι → M} (φ : M →SL[σ] M₂) {x : M} + (hf : has_sum f x) : + has_sum (λ (b:ι), φ (f b)) (φ x) := +by simpa only using hf.map φ.to_linear_map.to_add_monoid_hom φ.continuous + +alias continuous_linear_map.has_sum ← has_sum.mapL + +protected lemma continuous_linear_map.summable {f : ι → M} (φ : M →SL[σ] M₂) (hf : summable f) : + summable (λ b:ι, φ (f b)) := +(hf.has_sum.mapL φ).summable + +alias continuous_linear_map.summable ← summable.mapL + +protected lemma continuous_linear_map.map_tsum [t2_space M₂] {f : ι → M} + (φ : M →SL[σ] M₂) (hf : summable f) : φ (∑' z, f z) = ∑' z, φ (f z) := +(hf.has_sum.mapL φ).tsum_eq.symm + +include σ' +/-- Applying a continuous linear map commutes with taking an (infinite) sum. -/ +protected lemma continuous_linear_equiv.has_sum {f : ι → M} (e : M ≃SL[σ] M₂) {y : M₂} : + has_sum (λ (b:ι), e (f b)) y ↔ has_sum f (e.symm y) := +⟨λ h, by simpa only [e.symm.coe_coe, e.symm_apply_apply] using h.mapL (e.symm : M₂ →SL[σ'] M), + λ h, by simpa only [e.coe_coe, e.apply_symm_apply] using (e : M →SL[σ] M₂).has_sum h⟩ + +/-- Applying a continuous linear map commutes with taking an (infinite) sum. -/ +protected lemma continuous_linear_equiv.has_sum' {f : ι → M} (e : M ≃SL[σ] M₂) {x : M} : + has_sum (λ (b:ι), e (f b)) (e x) ↔ has_sum f x := +by rw [e.has_sum, continuous_linear_equiv.symm_apply_apply] + +protected lemma continuous_linear_equiv.summable {f : ι → M} (e : M ≃SL[σ] M₂) : + summable (λ b:ι, e (f b)) ↔ summable f := +⟨λ hf, (e.has_sum.1 hf.has_sum).summable, (e : M →SL[σ] M₂).summable⟩ + + +lemma continuous_linear_equiv.tsum_eq_iff [t2_space M] [t2_space M₂] {f : ι → M} + (e : M ≃SL[σ] M₂) {y : M₂} : ∑' z, e (f z) = y ↔ ∑' z, f z = e.symm y := +begin + by_cases hf : summable f, + { exact ⟨λ h, (e.has_sum.mp ((e.summable.mpr hf).has_sum_iff.mpr h)).tsum_eq, + λ h, (e.has_sum.mpr (hf.has_sum_iff.mpr h)).tsum_eq⟩ }, + { have hf' : ¬summable (λ z, e (f z)) := λ h, hf (e.summable.mp h), + rw [tsum_eq_zero_of_not_summable hf, tsum_eq_zero_of_not_summable hf'], + exact ⟨by { rintro rfl, simp }, λ H, by simpa using (congr_arg (λ z, e z) H)⟩ } +end + +protected lemma continuous_linear_equiv.map_tsum [t2_space M] [t2_space M₂] {f : ι → M} + (e : M ≃SL[σ] M₂) : e (∑' z, f z) = ∑' z, e (f z) := +by { refine symm (e.tsum_eq_iff.mpr _), rw e.symm_apply_apply _ } + +end has_sum From 56adee5b5eef9e734d82272918300fca4f3e7cef Mon Sep 17 00:00:00 2001 From: antoinelab01 Date: Wed, 18 Jan 2023 11:08:19 +0000 Subject: [PATCH 0280/1291] feat(combinatorics/quiver/single_obj): single object quivers (#17846) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit I think it makes sense for `category_theory/single_obj` to depend on `combinatorics/quiver/single_obj`, but I'd like to know what people working with category theory think of that. Co-authored by: Remi Bottinelli https://github.com/bottine Co-authored-by: antoinelab01 <66086247+antoinelab01@users.noreply.github.com> Co-authored-by: Antoine Labelle Co-authored-by: Rémi Bottinelli --- src/category_theory/single_obj.lean | 16 ++- src/combinatorics/quiver/basic.lean | 2 +- src/combinatorics/quiver/single_obj.lean | 121 +++++++++++++++++++++++ src/group_theory/nielsen_schreier.lean | 8 +- 4 files changed, 137 insertions(+), 10 deletions(-) create mode 100644 src/combinatorics/quiver/single_obj.lean diff --git a/src/category_theory/single_obj.lean b/src/category_theory/single_obj.lean index 4e7c97bd837d8..082a39641d817 100644 --- a/src/category_theory/single_obj.lean +++ b/src/category_theory/single_obj.lean @@ -6,6 +6,7 @@ Authors: Yury Kudryashov import category_theory.endomorphism import category_theory.category.Cat import algebra.category.Mon.basic +import combinatorics.quiver.single_obj /-! # Single-object category @@ -37,9 +38,11 @@ An element `x : α` can be reinterpreted as an element of `End (single_obj.star universes u v w namespace category_theory -/-- Type tag on `unit` used to define single-object categories and groupoids. -/ -@[nolint unused_arguments has_nonempty_instance] -def single_obj (α : Type u) : Type := unit + +/-- +Abbreviation that allows writing `category_theory.single_obj` rather than `quiver.single_obj`. +-/ +abbreviation single_obj := quiver.single_obj namespace single_obj @@ -75,8 +78,11 @@ instance groupoid [group α] : groupoid (single_obj α) := lemma inv_as_inv [group α] {x y : single_obj α} (f : x ⟶ y) : inv f = f⁻¹ := by { ext, rw [comp_as_mul, inv_mul_self, id_as_one] } -/-- The single object in `single_obj α`. -/ -protected def star : single_obj α := unit.star +/-- +Abbreviation that allows writing `category_theory.single_obj.star` rather than +`quiver.single_obj.star`. +-/ +abbreviation star : single_obj α := quiver.single_obj.star α /-- The endomorphisms monoid of the only object in `single_obj α` is equivalent to the original monoid α. -/ diff --git a/src/combinatorics/quiver/basic.lean b/src/combinatorics/quiver/basic.lean index 1dc41a020da10..245384adb31c4 100644 --- a/src/combinatorics/quiver/basic.lean +++ b/src/combinatorics/quiver/basic.lean @@ -103,7 +103,7 @@ lemma comp_assoc infix ` ⥤q `:50 := prefunctor -infix ` ⋙q `:50 := prefunctor.comp +infix ` ⋙q `:60 := prefunctor.comp notation `𝟭q` := id diff --git a/src/combinatorics/quiver/single_obj.lean b/src/combinatorics/quiver/single_obj.lean new file mode 100644 index 0000000000000..b77762c5b9fb8 --- /dev/null +++ b/src/combinatorics/quiver/single_obj.lean @@ -0,0 +1,121 @@ +/- +Copyright (c) 2023 Antoine Labelle. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Antoine Labelle +-/ +import combinatorics.quiver.cast +import combinatorics.quiver.symmetric + +/-! +# Single-object quiver + +Single object quiver with a given arrows type. + +## Main definitions + +Given a type `α`, `single_obj α` is the `unit` type, whose single object is called `star α`, with +`quiver` structure such that `star α ⟶ star α` is the type `α`. +An element `x : α` can be reinterpreted as an element of `star α ⟶ star α` using +`to_hom`. +More generally, a list of elements of `a` can be reinterpreted as a path from `star α` to +itself using `path_equiv_list`. +-/ + +namespace quiver + +/-- Type tag on `unit` used to define single-object quivers. -/ +@[derive unique, nolint unused_arguments] +def single_obj (α : Type*) : Type := unit + +namespace single_obj + +variables (α β γ : Type*) + +instance : quiver (single_obj α) := ⟨λ _ _, α⟩ + +/-- The single object in `single_obj α`. -/ +def star : single_obj α := unit.star + +instance : inhabited (single_obj α) := ⟨star α⟩ + +variables {α β γ} + +/-- Equip `single_obj α` with a reverse operation. -/ +@[reducible] -- See note [reducible non-instances] +def has_reverse (rev : α → α) : has_reverse (single_obj α) := ⟨λ _ _, rev⟩ + +/-- Equip `single_obj α` with an involutive reverse operation. -/ +@[reducible] -- See note [reducible non-instances] +def has_involutive_reverse (rev : α → α) (h : function.involutive rev) : + has_involutive_reverse (single_obj α) := +{ to_has_reverse := has_reverse rev, + inv' := λ _ _, h} + +/-- The type of arrows from `star α` to itself is equivalent to the original type `α`. -/ +@[simps] def to_hom : α ≃ (star α ⟶ star α) := equiv.refl _ + +/-- +Prefunctors between two `single_obj` quivers correspond to functions between the corresponding +arrows types. +-/ +@[simps] def to_prefunctor : + (α → β) ≃ (single_obj α ⥤q single_obj β) := +{ to_fun := λ f, ⟨id, λ _ _, f⟩, + inv_fun := λ f a, f.map (to_hom a), + left_inv := λ _, rfl, + right_inv := λ f, by cases f; obviously } + +lemma to_prefunctor_id : to_prefunctor id = 𝟭q (single_obj α) := rfl + +@[simp] lemma to_prefunctor_symm_id : + to_prefunctor.symm (𝟭q (single_obj α)) = id := rfl + +lemma to_prefunctor_comp (f : α → β) (g : β → γ) : + to_prefunctor (g ∘ f) = to_prefunctor f ⋙q to_prefunctor g := rfl + +@[simp] lemma to_prefunctor_symm_comp (f : single_obj α ⥤q single_obj β) + (g : single_obj β ⥤q single_obj γ) : to_prefunctor.symm (f ⋙q g) = + to_prefunctor.symm g ∘ to_prefunctor.symm f := +by simp only [equiv.symm_apply_eq, to_prefunctor_comp, equiv.apply_symm_apply] + +/-- +Auxiliary definition for `quiver.single_obj.path_equiv_list`. +Converts a path in the quiver `single_obj α` into a list of elements of type `a`. +-/ +@[simp] def path_to_list : Π {x : single_obj α}, path (star α) x → list α +| _ path.nil := [] +| _ (path.cons p a) := a :: path_to_list p + +/-- +Auxiliary definition for `quiver.single_obj.path_equiv_list`. +Converts a list of elements of type `α` into a path in the quiver `single_obj α`. +-/ +@[simp] def list_to_path : list α → path (star α) (star α) +| [] := path.nil +| (a :: l) := (list_to_path l).cons a + +lemma path_to_list_to_path {x : single_obj α} (p : path (star α) x) : + list_to_path (path_to_list p) = p.cast rfl (unit.ext) := +by { induction p with y z p a ih, refl, tidy, } + +lemma list_to_path_to_list (l : list α) : + path_to_list (list_to_path l) = l := +by { induction l with a l ih, refl, simp [ih] } + +/-- Paths in `single_obj α` quiver correspond to lists of elements of type `α`. -/ +def path_equiv_list : path (star α) (star α) ≃ list α := +⟨path_to_list, list_to_path, λ p, path_to_list_to_path p, list_to_path_to_list⟩ + +@[simp] lemma path_equiv_list_nil : path_equiv_list path.nil = ([] : list α) := rfl + +@[simp] lemma path_equiv_list_cons (p : path (star α) (star α)) (a : star α ⟶ star α) : + path_equiv_list (path.cons p a) = a :: path_to_list p := rfl + +@[simp] lemma path_equiv_list_symm_nil : path_equiv_list.symm ([] : list α) = path.nil := rfl + +@[simp] lemma path_equiv_list_symm_cons (l : list α) (a : α) : + path_equiv_list.symm (a :: l) = path.cons (path_equiv_list.symm l) a := rfl + +end single_obj + +end quiver diff --git a/src/group_theory/nielsen_schreier.lean b/src/group_theory/nielsen_schreier.lean index c36283d3646de..aa42a81781acd 100644 --- a/src/group_theory/nielsen_schreier.lean +++ b/src/group_theory/nielsen_schreier.lean @@ -71,7 +71,7 @@ class is_free_groupoid (G) [groupoid.{v} G] := (quiver_generators : quiver.{v+1} (is_free_groupoid.generators G)) (of : Π {a b : is_free_groupoid.generators G}, (a ⟶ b) → ((show G, from a) ⟶ b)) (unique_lift : ∀ {X : Type v} [group X] (f : labelling (is_free_groupoid.generators G) X), - ∃! F : G ⥤ single_obj X, ∀ a b (g : a ⟶ b), + ∃! F : G ⥤ category_theory.single_obj X, ∀ a b (g : a ⟶ b), F.map (of g) = f g) namespace is_free_groupoid @@ -82,7 +82,7 @@ attribute [instance] quiver_generators quiver. -/ @[ext] lemma ext_functor {G} [groupoid.{v} G] [is_free_groupoid G] {X : Type v} [group X] - (f g : G ⥤ single_obj X) + (f g : G ⥤ category_theory.single_obj X) (h : ∀ a b (e : a ⟶ b), f.map (of e) = g.map (of e)) : f = g := let ⟨_, _, u⟩ := @unique_lift G _ _ X _ (λ (a b : generators G) (e : a ⟶ b), g.map (of e)) in @@ -174,7 +174,7 @@ end /-- Since a hom gives a loop, any homomorphism from the vertex group at the root extends to a functor on the whole groupoid. -/ @[simps] def functor_of_monoid_hom {X} [monoid X] (f : End (root' T) →* X) : - G ⥤ single_obj X := + G ⥤ category_theory.single_obj X := { obj := λ _, (), map := λ a b p, f (loop_of_hom T p), map_id' := begin @@ -245,7 +245,7 @@ begin ←free_group.of_injective.eq_iff, ←mul_inv_eq_one], let X := free_group (weakly_connected_component $ generators G), let f : G → X := λ g, free_group.of (weakly_connected_component.mk g), - let F : G ⥤ single_obj X := single_obj.difference_functor f, + let F : G ⥤ category_theory.single_obj X := single_obj.difference_functor f, change F.map p = ((category_theory.functor.const G).obj ()).map p, congr, ext, rw [functor.const_obj_map, id_as_one, difference_functor_map, mul_inv_eq_one], From d1d69e99ed34c95266668af4e288fc1c598b9a7f Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Wed, 18 Jan 2023 12:53:10 +0000 Subject: [PATCH 0281/1291] feat(ring_theory/power_basis): matrix of multiplication by generator (#18177) + `power_basis.left_mul_matrix` computes the matrix of left multiplication by the generator with respect to the power basis: the last column is the negation of the coefficients of the minimal polynomial, while in other columns there are 1's below the diagonal and 0 elsewhere. + Also generalizes `comm_ring` or `field` in other files to `ring` where possible. --- .../dold_kan/normalized.lean | 11 ++- .../matrix/charpoly/minpoly.lean | 16 ++-- src/linear_algebra/matrix/to_lin.lean | 19 +++-- src/ring_theory/adjoin_root.lean | 4 +- src/ring_theory/norm.lean | 73 +++++++++---------- src/ring_theory/power_basis.lean | 18 +++++ 6 files changed, 81 insertions(+), 60 deletions(-) diff --git a/src/algebraic_topology/dold_kan/normalized.lean b/src/algebraic_topology/dold_kan/normalized.lean index d7d0db4ea424c..8a70211b6bf52 100644 --- a/src/algebraic_topology/dold_kan/normalized.lean +++ b/src/algebraic_topology/dold_kan/normalized.lean @@ -120,11 +120,18 @@ def N₁_iso_normalized_Moore_complex_comp_to_karoubi : { hom := { app := λ X, { f := P_infty_to_normalized_Moore_complex X, - comm := by tidy, }, }, + comm := by erw [comp_id, P_infty_comp_P_infty_to_normalized_Moore_complex] }, + naturality' := λ X Y f, by simp only [functor.comp_map, normalized_Moore_complex_map, + P_infty_to_normalized_Moore_complex_naturality, karoubi.hom_ext, karoubi.comp_f, N₁_map_f, + P_infty_comp_P_infty_to_normalized_Moore_complex_assoc, to_karoubi_map_f, assoc] }, inv := { app := λ X, { f := inclusion_of_Moore_complex_map X, - comm := by tidy, }, }, + comm := by erw [inclusion_of_Moore_complex_map_comp_P_infty, id_comp] }, + naturality' := λ X Y f, by { ext, simp only [functor.comp_map, normalized_Moore_complex_map, + karoubi.comp_f, to_karoubi_map_f, homological_complex.comp_f, normalized_Moore_complex.map_f, + inclusion_of_Moore_complex_map_f, factor_thru_arrow, N₁_map_f, + inclusion_of_Moore_complex_map_comp_P_infty_assoc, alternating_face_map_complex.map_f] } }, hom_inv_id' := begin ext X : 3, simp only [P_infty_to_normalized_Moore_complex_comp_inclusion_of_Moore_complex_map, diff --git a/src/linear_algebra/matrix/charpoly/minpoly.lean b/src/linear_algebra/matrix/charpoly/minpoly.lean index f0722405b0880..797e0cd6d565d 100644 --- a/src/linear_algebra/matrix/charpoly/minpoly.lean +++ b/src/linear_algebra/matrix/charpoly/minpoly.lean @@ -45,18 +45,16 @@ 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 := +lemma charpoly_left_mul_matrix {S : Type*} [ring S] [algebra R S] (h : power_basis R S) : + (left_mul_matrix h.basis h.gen).charpoly = minpoly R h.gen := begin - apply minpoly.unique, - { apply matrix.charpoly_monic }, + casesI subsingleton_or_nontrivial R, { apply subsingleton.elim }, + apply minpoly.unique' R h.gen (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 } + refine λ q hq, or_iff_not_imp_left.2 (λ h0, _), + rw [matrix.charpoly_degree_eq_dim, fintype.card_fin] at hq, + contrapose! hq, exact h.dim_le_degree_of_root h0 hq, end end power_basis diff --git a/src/linear_algebra/matrix/to_lin.lean b/src/linear_algebra/matrix/to_lin.lean index 45e7a98589889..7d5c8306db834 100644 --- a/src/linear_algebra/matrix/to_lin.lean +++ b/src/linear_algebra/matrix/to_lin.lean @@ -637,12 +637,8 @@ namespace algebra section lmul -variables {R S T : Type*} [comm_ring R] [comm_ring S] [comm_ring T] -variables [algebra R S] [algebra S T] [algebra R T] [is_scalar_tower R S T] -variables {m n : Type*} [fintype m] [decidable_eq m] [decidable_eq n] -variables (b : basis m R S) (c : basis n S T) - -open algebra +variables {R S : Type*} [comm_ring R] [ring S] [algebra R S] +variables {m : Type*} [fintype m] [decidable_eq m] (b : basis m R S) lemma to_matrix_lmul' (x : S) (i j) : linear_map.to_matrix b b (lmul R S x) i j = b.repr (x * b j) i := @@ -690,7 +686,14 @@ lemma left_mul_matrix_injective : function.injective (left_mul_matrix b) := ... = algebra.lmul R S x' 1 : by rw (linear_map.to_matrix b b).injective h ... = x' : mul_one x' -variable [fintype n] +end lmul + +section lmul_tower + +variables {R S T : Type*} [comm_ring R] [comm_ring S] [ring T] +variables [algebra R S] [algebra S T] [algebra R T] [is_scalar_tower R S T] +variables {m n : Type*} [fintype m] [fintype n] [decidable_eq m] [decidable_eq n] +variables (b : basis m R S) (c : basis n S T) lemma smul_left_mul_matrix (x) (ik jk) : left_mul_matrix (b.smul c) x ik jk = @@ -715,7 +718,7 @@ lemma smul_left_mul_matrix_algebra_map_ne (x : S) (i j) {k k'} (h : k ≠ k') : left_mul_matrix (b.smul c) (algebra_map _ _ x) (i, k) (j, k') = 0 := by rw [smul_left_mul_matrix_algebra_map, block_diagonal_apply_ne _ _ _ h] -end lmul +end lmul_tower end algebra diff --git a/src/ring_theory/adjoin_root.lean b/src/ring_theory/adjoin_root.lean index d0418ea5c80d7..4d8c81e2496b2 100644 --- a/src/ring_theory/adjoin_root.lean +++ b/src/ring_theory/adjoin_root.lean @@ -455,7 +455,7 @@ end end minpoly -section is_domain +section equiv' variables [comm_ring R] [comm_ring S] [algebra R S] variables (g : R[X]) (pb : _root_.power_basis R S) @@ -489,7 +489,7 @@ rfl (equiv' g pb h₁ h₂).symm.to_alg_hom = pb.lift (root g) h₁ := rfl -end is_domain +end equiv' section field diff --git a/src/ring_theory/norm.lean b/src/ring_theory/norm.lean index 9cd10eb83f041..f752917f45b59 100644 --- a/src/ring_theory/norm.lean +++ b/src/ring_theory/norm.lean @@ -41,7 +41,7 @@ See also `algebra.trace`, which is defined similarly as the trace of universes u v w -variables {R S T : Type*} [comm_ring R] [comm_ring S] +variables {R S T : Type*} [comm_ring R] [ring S] variables [algebra R S] variables {K L F : Type*} [field K] [field L] [field F] variables [algebra K L] [algebra K F] @@ -75,7 +75,7 @@ lemma norm_eq_matrix_det [fintype ι] [decidable_eq ι] (b : basis ι R S) (s : norm R s = matrix.det (algebra.left_mul_matrix b s) := by { rwa [norm_apply, ← linear_map.det_to_matrix b, ← to_matrix_lmul_eq], refl } -/-- If `x` is in the base field `K`, then the norm is `x ^ [L : K]`. -/ +/-- If `x` is in the base ring `K`, then the norm is `x ^ [L : K]`. -/ lemma norm_algebra_map_of_basis [fintype ι] (b : basis ι R S) (x : R) : norm R (algebra_map R S x) = x ^ fintype.card ι := begin @@ -91,7 +91,7 @@ end (If `L` is not finite-dimensional over `K`, then `norm = 1 = x ^ 0 = x ^ (finrank L K)`.) -/ @[simp] -protected lemma norm_algebra_map {K L : Type*} [field K] [comm_ring L] [algebra K L] (x : K) : +protected lemma norm_algebra_map {L : Type*} [ring L] [algebra K L] (x : K) : norm K (algebra_map K L x) = x ^ finrank K L := begin by_cases H : ∃ (s : finset L), nonempty (basis s K L), @@ -105,25 +105,24 @@ section eq_prod_roots /-- Given `pb : power_basis K S`, then the norm of `pb.gen` is `(-1) ^ pb.dim * coeff (minpoly K pb.gen) 0`. -/ -lemma power_basis.norm_gen_eq_coeff_zero_minpoly [algebra K S] (pb : power_basis K S) : - norm K pb.gen = (-1) ^ pb.dim * coeff (minpoly K pb.gen) 0 := -begin - rw [norm_eq_matrix_det pb.basis, det_eq_sign_charpoly_coeff, charpoly_left_mul_matrix, - fintype.card_fin] -end - -/-- Given `pb : power_basis K S`, then the norm of `pb.gen` is -`((minpoly K pb.gen).map (algebra_map K F)).roots.prod`. -/ -lemma power_basis.norm_gen_eq_prod_roots [algebra K S] (pb : power_basis K S) - (hf : (minpoly K pb.gen).splits (algebra_map K F)) : - algebra_map K F (norm K pb.gen) = - ((minpoly K pb.gen).map (algebra_map K F)).roots.prod := +lemma power_basis.norm_gen_eq_coeff_zero_minpoly (pb : power_basis R S) : + norm R pb.gen = (-1) ^ pb.dim * coeff (minpoly R pb.gen) 0 := +by rw [norm_eq_matrix_det pb.basis, det_eq_sign_charpoly_coeff, + charpoly_left_mul_matrix, fintype.card_fin] + +/-- Given `pb : power_basis R S`, then the norm of `pb.gen` is +`((minpoly R pb.gen).map (algebra_map R F)).roots.prod`. -/ +lemma power_basis.norm_gen_eq_prod_roots [algebra R F] (pb : power_basis R S) + (hf : (minpoly R pb.gen).splits (algebra_map R F)) : + algebra_map R F (norm R pb.gen) = + ((minpoly R pb.gen).map (algebra_map R F)).roots.prod := begin + haveI := module.nontrivial R F, + have := minpoly.monic pb.is_integral_gen, rw [power_basis.norm_gen_eq_coeff_zero_minpoly, ← pb.nat_degree_minpoly, ring_hom.map_mul, - ← coeff_map, prod_roots_eq_coeff_zero_of_monic_of_split - ((minpoly.monic (power_basis.is_integral_gen _)).map _) - ((splits_id_iff_splits _).2 hf), nat_degree_map, map_pow, ← mul_assoc, ← mul_pow], - simp + ← coeff_map, prod_roots_eq_coeff_zero_of_monic_of_split (this.map _) + ((splits_id_iff_splits _).2 hf), this.nat_degree_map, map_pow, ← mul_assoc, ← mul_pow], + { simp only [map_neg, _root_.map_one, neg_mul, neg_neg, one_pow, one_mul] }, apply_instance, end end eq_prod_roots @@ -225,20 +224,19 @@ section eq_prod_embeddings open intermediate_field intermediate_field.adjoin_simple polynomial -lemma norm_eq_prod_embeddings_gen {K L : Type*} [field K] [comm_ring L] [algebra K L] - (E : Type*) [field E] [algebra K E] - (pb : power_basis K L) - (hE : (minpoly K pb.gen).splits (algebra_map K E)) (hfx : (minpoly K pb.gen).separable) : - algebra_map K E (norm K pb.gen) = - (@@finset.univ (power_basis.alg_hom.fintype pb)).prod (λ σ, σ pb.gen) := +variables (F) (E : Type*) [field E] [algebra K E] + +lemma norm_eq_prod_embeddings_gen [algebra R F] (pb : power_basis R S) + (hE : (minpoly R pb.gen).splits (algebra_map R F)) (hfx : (minpoly R pb.gen).separable) : + algebra_map R F (norm R pb.gen) = (@@finset.univ pb^.alg_hom.fintype).prod (λ σ, σ pb.gen) := begin - letI := classical.dec_eq E, - rw [power_basis.norm_gen_eq_prod_roots pb hE, fintype.prod_equiv pb.lift_equiv', + letI := classical.dec_eq F, + rw [pb.norm_gen_eq_prod_roots hE, fintype.prod_equiv pb.lift_equiv', finset.prod_mem_multiset, finset.prod_eq_multiset_prod, multiset.to_finset_val, multiset.dedup_eq_self.mpr, multiset.map_id], - { exact nodup_roots ((separable_map _).mpr hfx) }, + { exact nodup_roots hfx.map }, { intro x, refl }, - { intro σ, rw [power_basis.lift_equiv'_apply_coe, id.def] } + { intro σ, rw [pb.lift_equiv'_apply_coe, id.def] } end lemma norm_eq_prod_roots [is_separable K L] [finite_dimensional K L] @@ -247,13 +245,10 @@ lemma norm_eq_prod_roots [is_separable K L] [finite_dimensional K L] by rw [norm_eq_norm_adjoin K x, map_pow, intermediate_field.adjoin_simple.norm_gen_eq_prod_roots _ hF] -variables (F) (E : Type*) [field E] [algebra K E] - lemma prod_embeddings_eq_finrank_pow [algebra L F] [is_scalar_tower K L F] [is_alg_closed E] [is_separable K F] [finite_dimensional K F] (pb : power_basis K L) : ∏ σ : F →ₐ[K] E, σ (algebra_map L F pb.gen) = - ((@@finset.univ (power_basis.alg_hom.fintype pb)).prod - (λ σ : L →ₐ[K] E, σ pb.gen)) ^ finrank L F := + ((@@finset.univ pb^.alg_hom.fintype).prod (λ σ : L →ₐ[K] E, σ pb.gen)) ^ finrank L F := begin haveI : finite_dimensional L F := finite_dimensional.right K L F, haveI : is_separable L F := is_separable_tower_top_of_is_separable K L F, @@ -301,18 +296,18 @@ begin ring_hom.id_apply] }, end -lemma is_integral_norm [algebra S L] [algebra S K] [is_scalar_tower S K L] - [is_separable K L] [finite_dimensional K L] {x : L} (hx : _root_.is_integral S x) : - _root_.is_integral S (norm K x) := +lemma is_integral_norm [algebra R L] [algebra R K] [is_scalar_tower R K L] + [is_separable K L] [finite_dimensional K L] {x : L} (hx : _root_.is_integral R x) : + _root_.is_integral R (norm K x) := begin have hx' : _root_.is_integral K x := is_integral_of_is_scalar_tower hx, rw [← is_integral_algebra_map_iff (algebra_map K (algebraic_closure L)).injective, norm_eq_prod_roots], { refine (is_integral.multiset_prod (λ y hy, _)).pow _, rw mem_roots_map (minpoly.ne_zero hx') at hy, - use [minpoly S x, minpoly.monic hx], + use [minpoly R x, minpoly.monic hx], rw ← aeval_def at ⊢ hy, - exact minpoly.aeval_of_is_scalar_tower S x y hy }, + exact minpoly.aeval_of_is_scalar_tower R x y hy }, { apply is_alg_closed.splits_codomain }, { apply_instance } end diff --git a/src/ring_theory/power_basis.lean b/src/ring_theory/power_basis.lean index 610f5a9153783..bece1cd0a14ef 100644 --- a/src/ring_theory/power_basis.lean +++ b/src/ring_theory/power_basis.lean @@ -214,6 +214,24 @@ lemma nat_degree_minpoly [nontrivial A] (pb : power_basis A S) : (minpoly A pb.gen).nat_degree = pb.dim := by rw [← minpoly_gen_eq, nat_degree_minpoly_gen] +protected lemma left_mul_matrix (pb : power_basis A S) : + algebra.left_mul_matrix pb.basis pb.gen = matrix.of + (λ i j, if ↑j + 1 = pb.dim then -pb.minpoly_gen.coeff ↑i else if ↑i = ↑j + 1 then 1 else 0) := +begin + casesI subsingleton_or_nontrivial A, { apply subsingleton.elim }, + rw [algebra.left_mul_matrix_apply, ← linear_equiv.eq_symm_apply, linear_map.to_matrix_symm], + refine pb.basis.ext (λ k, _), + simp_rw [matrix.to_lin_self, matrix.of_apply, pb.basis_eq_pow], + apply (pow_succ _ _).symm.trans, + split_ifs with h h, + { simp_rw [h, neg_smul, finset.sum_neg_distrib, eq_neg_iff_add_eq_zero], + convert pb.aeval_minpoly_gen, + rw [add_comm, aeval_eq_sum_range, finset.sum_range_succ, ← leading_coeff, + pb.minpoly_gen_monic.leading_coeff, one_smul, nat_degree_minpoly_gen, finset.sum_range] }, + { rw [fintype.sum_eq_single (⟨↑k + 1, lt_of_le_of_ne k.2 h⟩ : fin pb.dim), if_pos, one_smul], + { refl }, { refl }, intros x hx, rw [if_neg, zero_smul], apply mt fin.ext hx }, +end + end minpoly section equiv From fd838fdf07a83ca89fb66d30bebf6f0e02908c3f Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 18 Jan 2023 14:53:30 +0000 Subject: [PATCH 0282/1291] feat(data/list/of_fn): two lemmas about tuple concatenation (#18192) --- src/data/list/of_fn.lean | 9 +++++++++ 1 file changed, 9 insertions(+) diff --git a/src/data/list/of_fn.lean b/src/data/list/of_fn.lean index 18f556758be76..9a2410b3f0614 100644 --- a/src/data/list/of_fn.lean +++ b/src/data/list/of_fn.lean @@ -132,6 +132,10 @@ begin { rw [of_fn_succ', of_fn_succ', IH, append_concat], refl, }, end +@[simp] theorem of_fn_fin_append {m n} (a : fin m → α) (b : fin n → α) : + list.of_fn (fin.append a b) = list.of_fn a ++ list.of_fn b := +by simp_rw [of_fn_add, fin.append_left, fin.append_right] + /-- This breaks a list of `m*n` items into `m` groups each containing `n` elements. -/ theorem of_fn_mul {m n} (f : fin (m * n) → α) : list.of_fn f = list.join (list.of_fn $ λ i : fin m, list.of_fn $ λ j : fin n, @@ -173,6 +177,11 @@ by simp only [mem_of_fn, set.forall_range_iff] of_fn (λ i : fin n, c) = replicate n c := nat.rec_on n (by simp) $ λ n ihn, by simp [ihn] +@[simp] theorem of_fn_fin_repeat {m} (a : fin m → α) (n : ℕ) : + list.of_fn (fin.repeat n a) = (list.replicate n (list.of_fn a)).join := +by simp_rw [of_fn_mul, ←of_fn_const, fin.repeat, fin.mod_nat, fin.coe_mk, + add_comm, nat.add_mul_mod_self_right, nat.mod_eq_of_lt (fin.is_lt _), fin.eta] + /-- Lists are equivalent to the sigma type of tuples of a given length. -/ @[simps] def equiv_sigma_tuple : list α ≃ Σ n, fin n → α := From 7c3780f666ddb4ac9fb3b6d75a31c3e419d65973 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 18 Jan 2023 17:47:18 +0000 Subject: [PATCH 0283/1291] feat(algebra/triv_sq_zero_ext): lemmas about pow, and ring structure (#18199) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit `snd_pow : snd (x ^ n) = n • x.fst ^ n.pred • x.snd` captures the use of dual numbers in automatic differentiation. Also add myself to the authors, since I've done some work on this file in the past. --- src/algebra/triv_sq_zero_ext.lean | 85 +++++++++++++++++++++++++++++-- 1 file changed, 80 insertions(+), 5 deletions(-) diff --git a/src/algebra/triv_sq_zero_ext.lean b/src/algebra/triv_sq_zero_ext.lean index d04a32e95c5aa..fe789d9ff6354 100644 --- a/src/algebra/triv_sq_zero_ext.lean +++ b/src/algebra/triv_sq_zero_ext.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2020 Kenny Lau. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Kenny Lau +Authors: Kenny Lau, Eric Wieser -/ import algebra.algebra.basic @@ -76,12 +76,16 @@ section variables (M) @[simp] lemma fst_inl [has_zero M] (r : R) : (inl r : tsze R M).fst = r := rfl @[simp] lemma snd_inl [has_zero M] (r : R) : (inl r : tsze R M).snd = 0 := rfl +@[simp] lemma fst_comp_inl [has_zero M] : fst ∘ (inl : R → tsze R M) = id := rfl +@[simp] lemma snd_comp_inl [has_zero M] : snd ∘ (inl : R → tsze R M) = 0 := rfl end section variables (R) @[simp] lemma fst_inr [has_zero R] (m : M) : (inr m : tsze R M).fst = 0 := rfl @[simp] lemma snd_inr [has_zero R] (m : M) : (inr m : tsze R M).snd = m := rfl +@[simp] lemma fst_comp_inr [has_zero R] : fst ∘ (inr : M → tsze R M) = 0 := rfl +@[simp] lemma snd_comp_inr [has_zero R] : snd ∘ (inr : M → tsze R M) = id := rfl end lemma inl_injective [has_zero M] : function.injective (inl : R → tsze R M) := @@ -109,6 +113,9 @@ prod.has_zero instance [has_add R] [has_add M] : has_add (tsze R M) := prod.has_add +instance [has_sub R] [has_sub M] : has_sub (tsze R M) := +prod.has_sub + instance [has_neg R] [has_neg M] : has_neg (tsze R M) := prod.has_neg @@ -170,6 +177,11 @@ prod.module @[simp] lemma fst_neg [has_neg R] [has_neg M] (x : tsze R M) : (-x).fst = -x.fst := rfl @[simp] lemma snd_neg [has_neg R] [has_neg M] (x : tsze R M) : (-x).snd = -x.snd := rfl +@[simp] lemma fst_sub [has_sub R] [has_sub M] (x₁ x₂ : tsze R M) : + (x₁ - x₂).fst = x₁.fst - x₂.fst := rfl +@[simp] lemma snd_sub [has_sub R] [has_sub M] (x₁ x₂ : tsze R M) : + (x₁ - x₂).snd = x₁.snd - x₂.snd := rfl + @[simp] lemma fst_smul [has_smul S R] [has_smul S M] (s : S) (x : tsze R M) : (s • x).fst = s • x.fst := rfl @[simp] lemma snd_smul [has_smul S R] [has_smul S M] (s : S) (x : tsze R M) : @@ -184,10 +196,14 @@ variables (M) (inl (r₁ + r₂) : tsze R M) = inl r₁ + inl r₂ := ext rfl (add_zero 0).symm -@[simp] lemma inl_neg [has_neg R] [add_group M] (r : R) : +@[simp] lemma inl_neg [has_neg R] [sub_neg_zero_monoid M] (r : R) : (inl (-r) : tsze R M) = -inl r := ext rfl neg_zero.symm +@[simp] lemma inl_sub [has_sub R] [sub_neg_zero_monoid M] (r₁ r₂ : R) : + (inl (r₁ - r₂) : tsze R M) = inl r₁ - inl r₂ := +ext rfl (sub_zero _).symm + @[simp] lemma inl_smul [monoid S] [add_monoid M] [has_smul S R] [distrib_mul_action S M] (s : S) (r : R) : (inl (s • r) : tsze R M) = s • inl r := ext rfl (smul_zero s).symm @@ -203,10 +219,14 @@ variables (R) (inr (m₁ + m₂) : tsze R M) = inr m₁ + inr m₂ := ext (add_zero 0).symm rfl -@[simp] lemma inr_neg [add_group R] [has_neg M] (m : M) : +@[simp] lemma inr_neg [sub_neg_zero_monoid R] [has_neg M] (m : M) : (inr (-m) : tsze R M) = -inr m := ext neg_zero.symm rfl +@[simp] lemma inr_sub [sub_neg_zero_monoid R] [has_sub M] (m₁ m₂ : M) : + (inr (m₁ - m₂) : tsze R M) = inr m₁ - inr m₂ := +ext (sub_zero _).symm rfl + @[simp] lemma inr_smul [has_zero R] [has_zero S] [smul_with_zero S R] [has_smul S M] (r : S) (m : M) : (inr (r • m) : tsze R M) = r • inr m := ext (smul_zero _).symm rfl @@ -307,12 +327,33 @@ instance [monoid R] [add_monoid M] [distrib_mul_action R M] : mul_one_class (tsz .. triv_sq_zero_ext.has_mul } instance [add_monoid_with_one R] [add_monoid M] : add_monoid_with_one (tsze R M) := -{ nat_cast := λ n, (n, 0), +{ nat_cast := λ n, inl n, nat_cast_zero := by simp [nat.cast], nat_cast_succ := λ _, by ext; simp [nat.cast], .. triv_sq_zero_ext.add_monoid, .. triv_sq_zero_ext.has_one } +@[simp] lemma fst_nat_cast [add_monoid_with_one R] [add_monoid M] (n : ℕ) : + (n : tsze R M).fst = n := rfl +@[simp] lemma snd_nat_cast [add_monoid_with_one R] [add_monoid M] (n : ℕ) : + (n : tsze R M).snd = 0 := rfl +@[simp] lemma inl_nat_cast [add_monoid_with_one R] [add_monoid M] (n : ℕ) : + (inl n : tsze R M) = n := rfl + +instance [add_group_with_one R] [add_group M] : add_group_with_one (tsze R M) := +{ int_cast := λ z, inl z, + int_cast_of_nat := λ n, ext (int.cast_coe_nat _) rfl, + int_cast_neg_succ_of_nat := λ n, ext (int.cast_neg_succ_of_nat _) neg_zero.symm, + .. triv_sq_zero_ext.add_group, + .. triv_sq_zero_ext.add_monoid_with_one } + +@[simp] lemma fst_int_cast [add_group_with_one R] [add_group M] (z : ℤ) : + (z : tsze R M).fst = z := rfl +@[simp] lemma snd_int_cast [add_group_with_one R] [add_group M] (z : ℤ) : + (z : tsze R M).snd = 0 := rfl +@[simp] lemma inl_int_cast [add_group_with_one R] [add_group M] (z : ℤ) : + (inl z : tsze R M) = z := rfl + instance [semiring R] [add_comm_monoid M] [module R M] : non_assoc_semiring (tsze R M) := { zero_mul := λ x, ext (zero_mul x.1) $ show (0 : R) • x.2 + x.1 • 0 = 0, by rw [zero_smul, zero_add, smul_zero], @@ -330,11 +371,40 @@ instance [semiring R] [add_comm_monoid M] [module R M] : non_assoc_semiring (tsz .. triv_sq_zero_ext.mul_one_class, .. triv_sq_zero_ext.add_comm_monoid } +instance [ring R] [add_comm_group M] [module R M] : non_assoc_ring (tsze R M) := +{ .. triv_sq_zero_ext.add_group_with_one, + .. triv_sq_zero_ext.non_assoc_semiring } + +instance [comm_monoid R] [add_monoid M] [distrib_mul_action R M] : has_pow (tsze R M) ℕ := +⟨λ x n, ⟨x.fst^n, n • x.fst ^ n.pred • x.snd⟩⟩ + +@[simp] lemma fst_pow [comm_monoid R] [add_monoid M] [distrib_mul_action R M] + (x : tsze R M) (n : ℕ) : + fst (x ^ n) = x.fst ^ n := rfl + +@[simp] lemma snd_pow [comm_monoid R] [add_monoid M] [distrib_mul_action R M] + (x : tsze R M) (n : ℕ) : + snd (x ^ n) = n • x.fst ^ n.pred • x.snd := rfl + +@[simp] lemma inl_pow [comm_monoid R] [add_monoid M] [distrib_mul_action R M] + (r : R) (n : ℕ) : + (inl r ^ n : tsze R M) = inl (r ^ n) := +ext rfl $ by simp + instance [comm_monoid R] [add_monoid M] [distrib_mul_action R M] : monoid (tsze R M) := { mul_assoc := λ x y z, ext (mul_assoc x.1 y.1 z.1) $ show (x.1 * y.1) • z.2 + z.1 • (x.1 • y.2 + y.1 • x.2) = x.1 • (y.1 • z.2 + z.1 • y.2) + (y.1 * z.1) • x.2, by simp_rw [smul_add, ← mul_smul, add_assoc, mul_comm], + npow := λ n x, x ^ n, + npow_zero' := λ x, ext (pow_zero x.fst) (zero_smul _ _), + npow_succ' := λ n x, ext (pow_succ _ _) begin + dsimp, + rw [smul_comm (_ : R) n (_ : M), smul_smul, succ_nsmul'], + cases n, + { simp_rw [zero_smul ]}, + { rw [nat.pred_succ, pow_succ] } + end, .. triv_sq_zero_ext.mul_one_class } instance [comm_monoid R] [add_comm_monoid M] [distrib_mul_action R M] : comm_monoid (tsze R M) := @@ -346,6 +416,10 @@ instance [comm_semiring R] [add_comm_monoid M] [module R M] : comm_semiring (tsz { .. triv_sq_zero_ext.comm_monoid, .. triv_sq_zero_ext.non_assoc_semiring } +instance [comm_ring R] [add_comm_group M] [module R M] : comm_ring (tsze R M) := +{ .. triv_sq_zero_ext.non_assoc_ring, + .. triv_sq_zero_ext.comm_semiring } + variables (R M) /-- The canonical inclusion of rings `R → triv_sq_zero_ext R M`. -/ @@ -365,7 +439,8 @@ variables [comm_semiring S] [comm_semiring R] [add_comm_monoid M] variables [algebra S R] [module S M] [module R M] [is_scalar_tower S R M] instance algebra' : algebra S (tsze R M) := -{ commutes' := λ r x, mul_comm _ _, +{ smul := (•), + commutes' := λ r x, mul_comm _ _, smul_def' := λ r x, ext (algebra.smul_def _ _) $ show r • x.2 = algebra_map S R r • x.2 + x.1 • 0, by rw [smul_zero, add_zero, algebra_map_smul], .. (triv_sq_zero_ext.inl_hom R M).comp (algebra_map S R) } From 15a07a991d3429de3458649381c0e12d0c86c0b0 Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Wed, 18 Jan 2023 17:47:20 +0000 Subject: [PATCH 0284/1291] feat (analysis/integrals): integral of x * (1 + x^2)^t (#18201) --- src/analysis/special_functions/integrals.lean | 54 +++++++++++++++++++ 1 file changed, 54 insertions(+) diff --git a/src/analysis/special_functions/integrals.lean b/src/analysis/special_functions/integrals.lean index efd34ff0461cc..cedeb1c491154 100644 --- a/src/analysis/special_functions/integrals.lean +++ b/src/analysis/special_functions/integrals.lean @@ -400,6 +400,60 @@ end lemma integral_one_div_one_add_sq : ∫ x : ℝ in a..b, 1 / (1 + x^2) = arctan b - arctan a := by simp only [one_div, integral_inv_one_add_sq] +section rpow_cpow +open complex + +lemma integral_mul_cpow_one_add_sq {t : ℂ} (ht : t ≠ -1) : + ∫ x : ℝ in a..b, (x:ℂ) * (1 + x ^ 2) ^ t = + (1 + b ^ 2) ^ (t + 1) / (2 * (t + 1)) - (1 + a ^ 2) ^ (t + 1) / (2 * (t + 1)) := +begin + have : t + 1 ≠ 0 := by { contrapose! ht, rwa add_eq_zero_iff_eq_neg at ht }, + apply integral_eq_sub_of_has_deriv_at, + { intros x hx, + have f : has_deriv_at (λ y:ℂ, 1 + y ^ 2) (2 * x) x, + { convert (has_deriv_at_pow 2 (x:ℂ)).const_add 1, { norm_cast }, { simp } }, + have g : ∀ {z : ℂ}, (0 < z.re) → has_deriv_at (λ z, z ^ (t + 1) / (2 * (t + 1))) (z ^ t / 2) z, + { intros z hz, + have : z ≠ 0 := by { contrapose! hz, rw [hz, zero_re], }, + convert (has_deriv_at.cpow_const (has_deriv_at_id _) (or.inl hz)).div_const + (2 * (t + 1)) using 1, + field_simp, + ring }, + convert (has_deriv_at.comp ↑x (g _) f).comp_of_real using 1, + { field_simp, ring }, + { rw [add_re, one_re, ←of_real_pow, of_real_re], + exact (add_pos_of_pos_of_nonneg zero_lt_one (sq_nonneg x)) } }, + { apply continuous.interval_integrable, + refine continuous_of_real.mul _, + apply continuous.cpow, + { exact continuous_const.add (continuous_of_real.pow 2)}, + { exact continuous_const }, + { intro a, + rw [add_re, one_re, ←of_real_pow, of_real_re], + exact or.inl (add_pos_of_pos_of_nonneg zero_lt_one (sq_nonneg a)) } } +end + +lemma integral_mul_rpow_one_add_sq {t : ℝ} (ht : t ≠ -1) : + ∫ x : ℝ in a..b, x * (1 + x ^ 2) ^ t = + (1 + b ^ 2) ^ (t + 1) / (2 * (t + 1)) - (1 + a ^ 2) ^ (t + 1) / (2 * (t + 1)) := +begin + have : ∀ (x s : ℝ), (((1 + x ^ 2) ^ s : ℝ) : ℂ) = (1 + (x : ℂ) ^ 2) ^ ↑s, + { intros x s, + rw [of_real_cpow, of_real_add, of_real_pow, of_real_one], + exact add_nonneg zero_le_one (sq_nonneg x), }, + rw ←of_real_inj, + convert integral_mul_cpow_one_add_sq (_ : (t:ℂ) ≠ -1), + { rw ←interval_integral.integral_of_real, + congr' with x:1, + rw [of_real_mul, this x t] }, + { simp_rw [of_real_sub, of_real_div, this a (t + 1), this b (t + 1)], + push_cast }, + { rw [←of_real_one, ←of_real_neg, ne.def, of_real_inj], + exact ht }, +end + +end rpow_cpow + /-! ### Integral of `sin x ^ n` -/ lemma integral_sin_pow_aux : From cc70d9141824ea8982d1562ce009952f2c3ece30 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Wed, 18 Jan 2023 17:47:22 +0000 Subject: [PATCH 0285/1291] chore(*): add mathlib4 synchronization comments (#18202) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.gcd_monoid.finset` * `algebra.hom.centroid` * `combinatorics.set_family.compression.down` * `control.equiv_functor.instances` * `data.fin.tuple.basic` * `data.fin_enum` * `data.finset.lattice` * `data.finset.n_ary` * `data.finset.nat_antidiagonal` * `data.finset.pairwise` * `data.finset.powerset` * `data.finset.sigma` * `data.fintype.list` * `data.fintype.sigma` * `data.list.of_fn` * `data.matrix.dmatrix` * `data.real.basic` * `data.real.sign` * `group_theory.congruence` * `group_theory.perm.support` * `order.grade` * `order.hom.bounded` * `order.sup_indep` * `ring_theory.congruence` --- src/algebra/gcd_monoid/finset.lean | 3 +++ src/algebra/hom/centroid.lean | 3 +++ src/combinatorics/set_family/compression/down.lean | 3 +++ src/control/equiv_functor/instances.lean | 3 +++ src/data/fin/tuple/basic.lean | 3 +++ src/data/fin_enum.lean | 3 +++ src/data/finset/lattice.lean | 3 +++ src/data/finset/n_ary.lean | 3 +++ src/data/finset/nat_antidiagonal.lean | 3 +++ src/data/finset/pairwise.lean | 3 +++ src/data/finset/powerset.lean | 3 +++ src/data/finset/sigma.lean | 3 +++ src/data/fintype/list.lean | 3 +++ src/data/fintype/sigma.lean | 3 +++ src/data/list/of_fn.lean | 3 +++ src/data/matrix/dmatrix.lean | 3 +++ src/data/real/basic.lean | 3 +++ src/data/real/sign.lean | 3 +++ src/group_theory/congruence.lean | 3 +++ src/group_theory/perm/support.lean | 3 +++ src/order/grade.lean | 3 +++ src/order/hom/bounded.lean | 3 +++ src/order/sup_indep.lean | 3 +++ src/ring_theory/congruence.lean | 3 +++ 24 files changed, 72 insertions(+) diff --git a/src/algebra/gcd_monoid/finset.lean b/src/algebra/gcd_monoid/finset.lean index ddd10ea5a762b..b0857a36fe7ed 100644 --- a/src/algebra/gcd_monoid/finset.lean +++ b/src/algebra/gcd_monoid/finset.lean @@ -9,6 +9,9 @@ import algebra.gcd_monoid.multiset /-! # GCD and LCM operations on finsets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions - `finset.gcd` - the greatest common denominator of a `finset` of elements of a `gcd_monoid` diff --git a/src/algebra/hom/centroid.lean b/src/algebra/hom/centroid.lean index f074ac30ecd29..42578ad72ff4a 100644 --- a/src/algebra/hom/centroid.lean +++ b/src/algebra/hom/centroid.lean @@ -9,6 +9,9 @@ import algebra.hom.group_instances /-! # Centroid homomorphisms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Let `A` be a (non unital, non associative) algebra. The centroid of `A` is the set of linear maps `T` on `A` such that `T` commutes with left and right multiplication, that is to say, for all `a` and `b` in `A`, diff --git a/src/combinatorics/set_family/compression/down.lean b/src/combinatorics/set_family/compression/down.lean index 94d553f354236..2445f2eb75c10 100644 --- a/src/combinatorics/set_family/compression/down.lean +++ b/src/combinatorics/set_family/compression/down.lean @@ -8,6 +8,9 @@ import data.finset.card /-! # Down-compressions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines down-compression. Down-compressing `𝒜 : finset (finset α)` along `a : α` means removing `a` from the elements of `𝒜`, diff --git a/src/control/equiv_functor/instances.lean b/src/control/equiv_functor/instances.lean index f88b4ca3a9b37..793d7e910ae04 100644 --- a/src/control/equiv_functor/instances.lean +++ b/src/control/equiv_functor/instances.lean @@ -9,6 +9,9 @@ import control.equiv_functor /-! # `equiv_functor` instances +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We derive some `equiv_functor` instances, to enable `equiv_rw` to rewrite under these functions. -/ diff --git a/src/data/fin/tuple/basic.lean b/src/data/fin/tuple/basic.lean index b388460072fc0..d2b4cb189fac6 100644 --- a/src/data/fin/tuple/basic.lean +++ b/src/data/fin/tuple/basic.lean @@ -10,6 +10,9 @@ import data.set.intervals.basic /-! # Operation on tuples +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We interpret maps `Π i : fin n, α i` as `n`-tuples of elements of possibly varying type `α i`, `(α 0, …, α (n-1))`. A particular case is `fin n → α` of elements with all the same type. In this case when `α i` is a constant map, then tuples are isomorphic (but not definitionally equal) diff --git a/src/data/fin_enum.lean b/src/data/fin_enum.lean index 380f6a35cfebf..d8119788da8e0 100644 --- a/src/data/fin_enum.lean +++ b/src/data/fin_enum.lean @@ -8,6 +8,9 @@ import data.fintype.basic import data.list.prod_sigma /-! +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Type class for finitely enumerable types. The property is stronger than `fintype` in that it assigns each element a rank in a finite enumeration. diff --git a/src/data/finset/lattice.lean b/src/data/finset/lattice.lean index 321f110d5ab60..0a2be525db223 100644 --- a/src/data/finset/lattice.lean +++ b/src/data/finset/lattice.lean @@ -11,6 +11,9 @@ import order.complete_lattice /-! # Lattice operations on finsets + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ variables {α β γ ι : Type*} diff --git a/src/data/finset/n_ary.lean b/src/data/finset/n_ary.lean index 7b23ce0dd5208..34cef362469ca 100644 --- a/src/data/finset/n_ary.lean +++ b/src/data/finset/n_ary.lean @@ -8,6 +8,9 @@ import data.finset.prod /-! # N-ary images of finsets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + 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. diff --git a/src/data/finset/nat_antidiagonal.lean b/src/data/finset/nat_antidiagonal.lean index f5037e150821b..89b428effedb0 100644 --- a/src/data/finset/nat_antidiagonal.lean +++ b/src/data/finset/nat_antidiagonal.lean @@ -9,6 +9,9 @@ import data.multiset.nat_antidiagonal /-! # Antidiagonals in ℕ × ℕ as finsets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the antidiagonals of ℕ × ℕ as finsets: the `n`-th antidiagonal is the finset of pairs `(i, j)` such that `i + j = n`. This is useful for polynomial multiplication and more generally for sums going from `0` to `n`. diff --git a/src/data/finset/pairwise.lean b/src/data/finset/pairwise.lean index 846fd4d0ac2df..d71306e29b6c1 100644 --- a/src/data/finset/pairwise.lean +++ b/src/data/finset/pairwise.lean @@ -8,6 +8,9 @@ import data.finset.lattice /-! # Relations holding pairwise on finite sets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove a few results about the interaction of `set.pairwise_disjoint` and `finset`, as well as the interaction of `list.pairwise disjoint` and the condition of `disjoint` on `list.to_finset`, in `set` form. diff --git a/src/data/finset/powerset.lean b/src/data/finset/powerset.lean index 0dd929a6a0e97..adeeb63e6f1c4 100644 --- a/src/data/finset/powerset.lean +++ b/src/data/finset/powerset.lean @@ -8,6 +8,9 @@ import data.multiset.powerset /-! # The powerset of a finset + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ namespace finset diff --git a/src/data/finset/sigma.lean b/src/data/finset/sigma.lean index 6c463ce72a0fb..e000ed62e556e 100644 --- a/src/data/finset/sigma.lean +++ b/src/data/finset/sigma.lean @@ -9,6 +9,9 @@ import data.set.sigma /-! # Finite sets in a sigma type +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines a few `finset` constructions on `Σ i, α i`. ## Main declarations diff --git a/src/data/fintype/list.lean b/src/data/fintype/list.lean index 1accc08def6a3..2c9201bc56e31 100644 --- a/src/data/fintype/list.lean +++ b/src/data/fintype/list.lean @@ -10,6 +10,9 @@ import data.finset.powerset # Fintype instance for nodup lists +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The subtype of `{l : list α // l.nodup}` over a `[fintype α]` admits a `fintype` instance. diff --git a/src/data/fintype/sigma.lean b/src/data/fintype/sigma.lean index 17ca0c138b297..7cea94a731f50 100644 --- a/src/data/fintype/sigma.lean +++ b/src/data/fintype/sigma.lean @@ -9,6 +9,9 @@ import data.finset.sigma /-! # fintype instances for sigma types + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ open function diff --git a/src/data/list/of_fn.lean b/src/data/list/of_fn.lean index 9a2410b3f0614..4a578e41e8680 100644 --- a/src/data/list/of_fn.lean +++ b/src/data/list/of_fn.lean @@ -10,6 +10,9 @@ import data.list.join /-! # Lists from functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Theorems and lemmas for dealing with `list.of_fn`, which converts a function on `fin n` to a list of length `n`. diff --git a/src/data/matrix/dmatrix.lean b/src/data/matrix/dmatrix.lean index 2fe70c3422ed0..4a80dde055876 100644 --- a/src/data/matrix/dmatrix.lean +++ b/src/data/matrix/dmatrix.lean @@ -8,6 +8,9 @@ import data.fintype.basic /-! # Matrices + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ universes u u' v w z diff --git a/src/data/real/basic.lean b/src/data/real/basic.lean index 76aa36a70ac09..4009cb9258435 100644 --- a/src/data/real/basic.lean +++ b/src/data/real/basic.lean @@ -11,6 +11,9 @@ import data.real.cau_seq_completion /-! # Real numbers from Cauchy sequences +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines `ℝ` as the type of equivalence classes of Cauchy sequences of rational numbers. This choice is motivated by how easy it is to prove that `ℝ` is a commutative ring, by simply lifting everything to `ℚ`. diff --git a/src/data/real/sign.lean b/src/data/real/sign.lean index eca1c6d1892b1..34f156eb1119a 100644 --- a/src/data/real/sign.lean +++ b/src/data/real/sign.lean @@ -8,6 +8,9 @@ import data.real.basic /-! # Real sign function +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file introduces and contains some results about `real.sign` which maps negative real numbers to -1, positive real numbers to 1, and 0 to 0. diff --git a/src/group_theory/congruence.lean b/src/group_theory/congruence.lean index 11748fd953fb4..857995ec74330 100644 --- a/src/group_theory/congruence.lean +++ b/src/group_theory/congruence.lean @@ -11,6 +11,9 @@ import group_theory.submonoid.operations /-! # Congruence relations +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines congruence relations: equivalence relations that preserve a binary operation, which in this case is multiplication or addition. The principal definition is a `structure` extending a `setoid` (an equivalence relation), and the inductive definition of the smallest diff --git a/src/group_theory/perm/support.lean b/src/group_theory/perm/support.lean index f8aac0b2bf1e7..78fa4c4dc079a 100644 --- a/src/group_theory/perm/support.lean +++ b/src/group_theory/perm/support.lean @@ -10,6 +10,9 @@ import group_theory.perm.basic /-! # Support of a permutation +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions In the following, `f g : equiv.perm α`. diff --git a/src/order/grade.lean b/src/order/grade.lean index 9d693b59419ca..1f0c44cac0b23 100644 --- a/src/order/grade.lean +++ b/src/order/grade.lean @@ -9,6 +9,9 @@ import data.int.succ_pred /-! # Graded orders +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines graded orders, also known as ranked orders. A `𝕆`-graded order is an order `α` equipped with a distinguished "grade" function `α → 𝕆` which diff --git a/src/order/hom/bounded.lean b/src/order/hom/bounded.lean index 359ab4bfd14d0..cb58486e39e07 100644 --- a/src/order/hom/bounded.lean +++ b/src/order/hom/bounded.lean @@ -9,6 +9,9 @@ import order.bounded_order /-! # Bounded order homomorphisms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines (bounded) order homomorphisms. We use the `fun_like` design, so each type of morphisms has a companion typeclass which is meant to diff --git a/src/order/sup_indep.lean b/src/order/sup_indep.lean index 70c982f5ccad4..487c9e623553d 100644 --- a/src/order/sup_indep.lean +++ b/src/order/sup_indep.lean @@ -10,6 +10,9 @@ import data.fintype.basic /-! # Supremum independence +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we define supremum independence of indexed sets. An indexed family `f : ι → α` is sup-independent if, for all `a`, `f a` and the supremum of the rest are disjoint. diff --git a/src/ring_theory/congruence.lean b/src/ring_theory/congruence.lean index a79ed4a5a51b9..f0c339ea28545 100644 --- a/src/ring_theory/congruence.lean +++ b/src/ring_theory/congruence.lean @@ -12,6 +12,9 @@ import group_theory.congruence /-! # Congruence relations on rings +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines congruence relations on rings, which extend `con` and `add_con` on monoids and additive monoids. From 97586a95edd9e36adfc786788500bce489ccacfa Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Wed, 18 Jan 2023 17:47:23 +0000 Subject: [PATCH 0286/1291] chore(ring_theory/adjoin_root): fix timeout by writing out `simps` manually (#18209) The `@[simps]` attribute on `adjoin_root.power_basis_aux'` is prone to timing out. Since `adjoin_root.power_basis_aux'` is defined as just a constructor application, we can't easily fix this by forbidding it from reducing or simplifying the result. So we'll have to work around this issue by defining the `@[simp]` lemmas manually. I left out `adjoin_root.power_basis_aux'_repr_apply_support_val` which says something about casting the support of the `repr` function to a multiset, since I can't imagine that one ever being useful, and we should be able to prove it for all `basis.of_equiv_fun` definitions anyway. --- src/ring_theory/adjoin_root.lean | 13 +++++++++++-- 1 file changed, 11 insertions(+), 2 deletions(-) diff --git a/src/ring_theory/adjoin_root.lean b/src/ring_theory/adjoin_root.lean index 4d8c81e2496b2..6fb17d1e38b18 100644 --- a/src/ring_theory/adjoin_root.lean +++ b/src/ring_theory/adjoin_root.lean @@ -334,8 +334,17 @@ basis.of_equiv_fun exact (degree_eq_nat_degree $ hg.ne_zero).symm ▸ degree_sum_fin_lt _ }, end} --- This was moved after the definition to prevent a timeout -attribute [simps] power_basis_aux' +/-- This lemma could be autogenerated by `@[simps]` but unfortunately that would require +unfolding that causes a timeout. -/ +@[simp] lemma power_basis_aux'_repr_symm_apply (hg : g.monic) (c : fin g.nat_degree →₀ R) : + (power_basis_aux' hg).repr.symm c = mk g (∑ (i : fin _), monomial i (c i)) := rfl + +/-- This lemma could be autogenerated by `@[simps]` but unfortunately that would require +unfolding that causes a timeout. -/ +@[simp] theorem power_basis_aux'_repr_apply_to_fun (hg : g.monic) (f : adjoin_root g) + (i : fin g.nat_degree) : + (power_basis_aux' hg).repr f i = (mod_by_monic_hom hg f).coeff ↑i := +rfl /-- The power basis `1, root g, ..., root g ^ (d - 1)` for `adjoin_root g`, where `g` is a monic polynomial of degree `d`. -/ From df78eae582aad2f545024bf6c7249191d2723074 Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Wed, 18 Jan 2023 17:47:25 +0000 Subject: [PATCH 0287/1291] refactor(geometry/euclidean/angle/unoriented/basic): split out conformal map lemmas (#18210) Move two lemmas about conformal maps preserving angles to a separate `geometry.euclidean.angle.unoriented.conformal`, so eliminating the dependence of angles on calculus (and use `assert_not_exists` to declare that the basic file should not depend on calculus or conformal maps). None of the other Euclidean geometry results depend on these two lemmas. --- .../calculus/conformal/normed_space.lean | 2 +- .../euclidean/angle/unoriented/basic.lean | 22 ++--------- .../euclidean/angle/unoriented/conformal.lean | 38 +++++++++++++++++++ 3 files changed, 42 insertions(+), 20 deletions(-) create mode 100644 src/geometry/euclidean/angle/unoriented/conformal.lean diff --git a/src/analysis/calculus/conformal/normed_space.lean b/src/analysis/calculus/conformal/normed_space.lean index 5fdcd41193683..03ad42c84f295 100644 --- a/src/analysis/calculus/conformal/normed_space.lean +++ b/src/analysis/calculus/conformal/normed_space.lean @@ -26,7 +26,7 @@ if it is real differentiable at that point and its differential `is_conformal_li In `analysis.calculus.conformal.inner_product`: * `conformal_at_iff`: an equivalent definition of the conformality of a map -In `geometry.euclidean.basic`: +In `geometry.euclidean.angle.unoriented.conformal`: * `conformal_at.preserves_angle`: if a map is conformal at `x`, then its differential preserves all angles at `x` diff --git a/src/geometry/euclidean/angle/unoriented/basic.lean b/src/geometry/euclidean/angle/unoriented/basic.lean index b361b70d0be13..6afbd85c4ec5c 100644 --- a/src/geometry/euclidean/angle/unoriented/basic.lean +++ b/src/geometry/euclidean/angle/unoriented/basic.lean @@ -3,7 +3,6 @@ Copyright (c) 2020 Joseph Myers. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Joseph Myers, Manuel Candales -/ -import analysis.calculus.conformal.normed_space import analysis.inner_product_space.basic import analysis.special_functions.trigonometric.inverse @@ -18,6 +17,9 @@ This file defines unoriented angles in real inner product spaces. -/ +assert_not_exists has_fderiv_at +assert_not_exists conformal_at + noncomputable theory open_locale big_operators open_locale real @@ -53,24 +55,6 @@ by rw [angle, angle, f.inner_map_map, f.norm_map, f.norm_map] angle (x : V) (y : V) = angle x y := s.subtypeₗᵢ.angle_map x y -lemma is_conformal_map.preserves_angle {E F : Type*} - [inner_product_space ℝ E] [inner_product_space ℝ F] - {f' : E →L[ℝ] F} (h : is_conformal_map f') (u v : E) : - angle (f' u) (f' v) = angle u v := -begin - obtain ⟨c, hc, li, rfl⟩ := h, - exact (angle_smul_smul hc _ _).trans (li.angle_map _ _) -end - -/-- If a real differentiable map `f` is conformal at a point `x`, - then it preserves the angles at that point. -/ -lemma conformal_at.preserves_angle {E F : Type*} - [inner_product_space ℝ E] [inner_product_space ℝ F] - {f : E → F} {x : E} {f' : E →L[ℝ] F} - (h : has_fderiv_at f f' x) (H : conformal_at f x) (u v : E) : - angle (f' u) (f' v) = angle u v := -let ⟨f₁, h₁, c⟩ := H in h₁.unique h ▸ is_conformal_map.preserves_angle c u v - /-- The cosine of the angle between two vectors. -/ lemma cos_angle (x y : V) : real.cos (angle x y) = ⟪x, y⟫ / (‖x‖ * ‖y‖) := real.cos_arccos (abs_le.mp (abs_real_inner_div_norm_mul_norm_le_one x y)).1 diff --git a/src/geometry/euclidean/angle/unoriented/conformal.lean b/src/geometry/euclidean/angle/unoriented/conformal.lean new file mode 100644 index 0000000000000..cd65fa3ca0463 --- /dev/null +++ b/src/geometry/euclidean/angle/unoriented/conformal.lean @@ -0,0 +1,38 @@ +/- +Copyright (c) 2021 Yourong Zang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yourong Zang +-/ +import analysis.calculus.conformal.normed_space +import geometry.euclidean.angle.unoriented.basic + +/-! +# Angles and conformal maps + +This file proves that conformal maps preserve angles. + +-/ + +namespace inner_product_geometry + +variables {V : Type*} [inner_product_space ℝ V] + +lemma is_conformal_map.preserves_angle {E F : Type*} + [inner_product_space ℝ E] [inner_product_space ℝ F] + {f' : E →L[ℝ] F} (h : is_conformal_map f') (u v : E) : + angle (f' u) (f' v) = angle u v := +begin + obtain ⟨c, hc, li, rfl⟩ := h, + exact (angle_smul_smul hc _ _).trans (li.angle_map _ _) +end + +/-- If a real differentiable map `f` is conformal at a point `x`, + then it preserves the angles at that point. -/ +lemma conformal_at.preserves_angle {E F : Type*} + [inner_product_space ℝ E] [inner_product_space ℝ F] + {f : E → F} {x : E} {f' : E →L[ℝ] F} + (h : has_fderiv_at f f' x) (H : conformal_at f x) (u v : E) : + angle (f' u) (f' v) = angle u v := +let ⟨f₁, h₁, c⟩ := H in h₁.unique h ▸ is_conformal_map.preserves_angle c u v + +end inner_product_geometry From fb319896dcaa409bd4a3cc0f8484297ef9dae2c3 Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Wed, 18 Jan 2023 17:47:27 +0000 Subject: [PATCH 0288/1291] refactor(geometry/euclidean/angle/oriented/basic): split out rotations (#18212) `geometry.euclidean.angle.oriented.basic` is 1489 lines long. Reduce it to 1032 lines by splitting out the definition of and results about `rotation` to a separate file. (No results in this file involve `rotation` in their proofs unless it's also involved in the statement.) There are no changes to API or proofs. --- .../euclidean/angle/oriented/affine.lean | 2 +- .../euclidean/angle/oriented/basic.lean | 457 ---------------- .../euclidean/angle/oriented/rotation.lean | 488 ++++++++++++++++++ 3 files changed, 489 insertions(+), 458 deletions(-) create mode 100644 src/geometry/euclidean/angle/oriented/rotation.lean diff --git a/src/geometry/euclidean/angle/oriented/affine.lean b/src/geometry/euclidean/angle/oriented/affine.lean index c627fc75e2991..ccb652414fe37 100644 --- a/src/geometry/euclidean/angle/oriented/affine.lean +++ b/src/geometry/euclidean/angle/oriented/affine.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Joseph Myers -/ import analysis.convex.side -import geometry.euclidean.angle.oriented.basic +import geometry.euclidean.angle.oriented.rotation import geometry.euclidean.angle.unoriented.affine /-! diff --git a/src/geometry/euclidean/angle/oriented/basic.lean b/src/geometry/euclidean/angle/oriented/basic.lean index 8f1c052d56586..c4d739c33fcc8 100644 --- a/src/geometry/euclidean/angle/oriented/basic.lean +++ b/src/geometry/euclidean/angle/oriented/basic.lean @@ -16,8 +16,6 @@ This file defines oriented angles in real inner product spaces. * `orientation.oangle` is the oriented angle between two vectors with respect to an orientation. -* `orientation.rotation` is the rotation by an oriented angle with respect to an orientation. - ## Implementation notes The definitions here use the `real.angle` type, angles modulo `2 * π`. For some purposes, @@ -45,7 +43,6 @@ variables {V V' : Type*} [inner_product_space ℝ V] [inner_product_space ℝ V' variables [fact (finrank ℝ V = 2)] [fact (finrank ℝ V' = 2)] (o : orientation ℝ V (fin 2)) local notation `ω` := o.area_form -local notation `J` := o.right_angle_rotation /-- The oriented angle from `x` to `y`, modulo `2 * π`. If either vector is 0, this is 0. See `inner_product_geometry.angle` for the corresponding unoriented angle definition. -/ @@ -553,357 +550,6 @@ begin simp end -/-- Auxiliary construction to build a rotation by the oriented angle `θ`. -/ -def rotation_aux (θ : real.angle) : V →ₗᵢ[ℝ] V := -linear_map.isometry_of_inner - (real.angle.cos θ • linear_map.id - + real.angle.sin θ • ↑(linear_isometry_equiv.to_linear_equiv J)) - begin - intros x y, - simp only [is_R_or_C.conj_to_real, id.def, linear_map.smul_apply, linear_map.add_apply, - linear_map.id_coe, linear_equiv.coe_coe, linear_isometry_equiv.coe_to_linear_equiv, - orientation.area_form_right_angle_rotation_left, - orientation.inner_right_angle_rotation_left, - orientation.inner_right_angle_rotation_right, - inner_add_left, inner_smul_left, inner_add_right, inner_smul_right], - linear_combination inner x y * θ.cos_sq_add_sin_sq, - end - -@[simp] lemma rotation_aux_apply (θ : real.angle) (x : V) : - o.rotation_aux θ x = real.angle.cos θ • x + real.angle.sin θ • J x := -rfl - -/-- A rotation by the oriented angle `θ`. -/ -def rotation (θ : real.angle) : V ≃ₗᵢ[ℝ] V := -linear_isometry_equiv.of_linear_isometry - (o.rotation_aux θ) - (real.angle.cos θ • linear_map.id - real.angle.sin θ • ↑(linear_isometry_equiv.to_linear_equiv J)) - begin - ext x, - convert congr_arg (λ t : ℝ, t • x) θ.cos_sq_add_sin_sq using 1, - { simp only [o.right_angle_rotation_right_angle_rotation, o.rotation_aux_apply, - function.comp_app, id.def, linear_equiv.coe_coe, linear_isometry.coe_to_linear_map, - linear_isometry_equiv.coe_to_linear_equiv, map_smul, map_sub, linear_map.coe_comp, - linear_map.id_coe, linear_map.smul_apply, linear_map.sub_apply, ← mul_smul, add_smul, - smul_add, smul_neg, smul_sub, mul_comm, sq], - abel }, - { simp }, - end - begin - ext x, - convert congr_arg (λ t : ℝ, t • x) θ.cos_sq_add_sin_sq using 1, - { simp only [o.right_angle_rotation_right_angle_rotation, o.rotation_aux_apply, - function.comp_app, id.def, linear_equiv.coe_coe, linear_isometry.coe_to_linear_map, - linear_isometry_equiv.coe_to_linear_equiv, map_add, map_smul, linear_map.coe_comp, - linear_map.id_coe, linear_map.smul_apply, linear_map.sub_apply, add_smul, ← mul_smul, - mul_comm, smul_add, smul_neg, sq], - abel }, - { simp }, - end - -lemma rotation_apply (θ : real.angle) (x : V) : - o.rotation θ x = real.angle.cos θ • x + real.angle.sin θ • J x := -rfl - -lemma rotation_symm_apply (θ : real.angle) (x : V) : - (o.rotation θ).symm x = real.angle.cos θ • x - real.angle.sin θ • J x := -rfl - -attribute [irreducible] rotation - -lemma rotation_eq_matrix_to_lin (θ : real.angle) {x : V} (hx : x ≠ 0) : - (o.rotation θ).to_linear_map - = matrix.to_lin - (o.basis_right_angle_rotation x hx) (o.basis_right_angle_rotation x hx) - !![θ.cos, -θ.sin; θ.sin, θ.cos] := -begin - apply (o.basis_right_angle_rotation x hx).ext, - intros i, - fin_cases i, - { rw matrix.to_lin_self, - simp [rotation_apply, fin.sum_univ_succ] }, - { rw matrix.to_lin_self, - simp [rotation_apply, fin.sum_univ_succ, add_comm] }, -end - -/-- The determinant of `rotation` (as a linear map) is equal to `1`. -/ -@[simp] lemma det_rotation (θ : real.angle) : - (o.rotation θ).to_linear_map.det = 1 := -begin - haveI : nontrivial V := - finite_dimensional.nontrivial_of_finrank_eq_succ (fact.out (finrank ℝ V = 2)), - obtain ⟨x, hx⟩ : ∃ x, x ≠ (0:V) := exists_ne (0:V), - rw o.rotation_eq_matrix_to_lin θ hx, - simpa [sq] using θ.cos_sq_add_sin_sq, -end - -/-- The determinant of `rotation` (as a linear equiv) is equal to `1`. -/ -@[simp] lemma linear_equiv_det_rotation (θ : real.angle) : - (o.rotation θ).to_linear_equiv.det = 1 := -units.ext $ o.det_rotation θ - -/-- The inverse of `rotation` is rotation by the negation of the angle. -/ -@[simp] lemma rotation_symm (θ : real.angle) : (o.rotation θ).symm = o.rotation (-θ) := -by ext; simp [o.rotation_apply, o.rotation_symm_apply, sub_eq_add_neg] - -/-- Rotation by 0 is the identity. -/ -@[simp] lemma rotation_zero : o.rotation 0 = linear_isometry_equiv.refl ℝ V := -by ext; simp [rotation] - -/-- Rotation by π is negation. -/ -@[simp] lemma rotation_pi : o.rotation π = linear_isometry_equiv.neg ℝ := -begin - ext x, - simp [rotation] -end - -/-- Rotation by π is negation. -/ -lemma rotation_pi_apply (x : V) : o.rotation π x = -x := -by simp - -/-- Rotation by π / 2 is the "right-angle-rotation" map `J`. -/ -lemma rotation_pi_div_two : o.rotation (π / 2 : ℝ) = J := -begin - ext x, - simp [rotation], -end - -/-- Rotating twice is equivalent to rotating by the sum of the angles. -/ -@[simp] lemma rotation_rotation (θ₁ θ₂ : real.angle) (x : V) : - o.rotation θ₁ (o.rotation θ₂ x) = o.rotation (θ₁ + θ₂) x := -begin - simp only [o.rotation_apply, ←mul_smul, real.angle.cos_add, real.angle.sin_add, add_smul, - sub_smul, linear_isometry_equiv.trans_apply, smul_add, linear_isometry_equiv.map_add, - linear_isometry_equiv.map_smul, right_angle_rotation_right_angle_rotation, smul_neg], - ring_nf, - abel, -end - -/-- Rotating twice is equivalent to rotating by the sum of the angles. -/ -@[simp] lemma rotation_trans (θ₁ θ₂ : real.angle) : - (o.rotation θ₁).trans (o.rotation θ₂) = o.rotation (θ₂ + θ₁) := -linear_isometry_equiv.ext $ λ _, by rw [←rotation_rotation, linear_isometry_equiv.trans_apply] - -/-- Rotating the first of two vectors by `θ` scales their Kahler form by `cos θ - sin θ * I`. -/ -@[simp] lemma kahler_rotation_left (x y : V) (θ : real.angle) : - o.kahler (o.rotation θ x) y = conj (θ.exp_map_circle : ℂ) * o.kahler x y := -begin - simp only [o.rotation_apply, map_add, map_mul, linear_map.map_smulₛₗ, ring_hom.id_apply, - linear_map.add_apply, linear_map.smul_apply, real_smul, kahler_right_angle_rotation_left, - real.angle.coe_exp_map_circle, is_R_or_C.conj_of_real, conj_I], - ring, -end - -/-- Negating a rotation is equivalent to rotation by π plus the angle. -/ -lemma neg_rotation (θ : real.angle) (x : V) : -o.rotation θ x = o.rotation (π + θ) x := -by rw [←o.rotation_pi_apply, rotation_rotation] - -/-- Negating a rotation by -π / 2 is equivalent to rotation by π / 2. -/ -@[simp] lemma neg_rotation_neg_pi_div_two (x : V) : - -o.rotation (-π / 2 : ℝ) x = o.rotation (π / 2 : ℝ) x := -by rw [neg_rotation, ←real.angle.coe_add, neg_div, ←sub_eq_add_neg, sub_half] - -/-- Negating a rotation by π / 2 is equivalent to rotation by -π / 2. -/ -lemma neg_rotation_pi_div_two (x : V) : -o.rotation (π / 2 : ℝ) x = o.rotation (-π / 2 : ℝ) x := -neg_eq_iff_neg_eq.1 $ o.neg_rotation_neg_pi_div_two _ - -/-- Rotating the first of two vectors by `θ` scales their Kahler form by `cos (-θ) + sin (-θ) * I`. --/ -lemma kahler_rotation_left' (x y : V) (θ : real.angle) : - o.kahler (o.rotation θ x) y = (-θ).exp_map_circle * o.kahler x y := -by simpa [coe_inv_circle_eq_conj, -kahler_rotation_left] using o.kahler_rotation_left x y θ - -/-- Rotating the second of two vectors by `θ` scales their Kahler form by `cos θ + sin θ * I`. -/ -@[simp] lemma kahler_rotation_right (x y : V) (θ : real.angle) : - o.kahler x (o.rotation θ y) = θ.exp_map_circle * o.kahler x y := -begin - simp only [o.rotation_apply, map_add, linear_map.map_smulₛₗ, ring_hom.id_apply, real_smul, - kahler_right_angle_rotation_right, real.angle.coe_exp_map_circle], - ring, -end - -/-- Rotating the first vector by `θ` subtracts `θ` from the angle between two vectors. -/ -@[simp] lemma oangle_rotation_left {x y : V} (hx : x ≠ 0) (hy : y ≠ 0) (θ : real.angle) : - o.oangle (o.rotation θ x) y = o.oangle x y - θ := -begin - simp only [oangle, o.kahler_rotation_left'], - rw [complex.arg_mul_coe_angle, real.angle.arg_exp_map_circle], - { abel }, - { exact ne_zero_of_mem_circle _ }, - { exact o.kahler_ne_zero hx hy }, -end - -/-- Rotating the second vector by `θ` adds `θ` to the angle between two vectors. -/ -@[simp] lemma oangle_rotation_right {x y : V} (hx : x ≠ 0) (hy : y ≠ 0) (θ : real.angle) : - o.oangle x (o.rotation θ y) = o.oangle x y + θ := -begin - simp only [oangle, o.kahler_rotation_right], - rw [complex.arg_mul_coe_angle, real.angle.arg_exp_map_circle], - { abel }, - { exact ne_zero_of_mem_circle _ }, - { exact o.kahler_ne_zero hx hy }, -end - -/-- The rotation of a vector by `θ` has an angle of `-θ` from that vector. -/ -@[simp] lemma oangle_rotation_self_left {x : V} (hx : x ≠ 0) (θ : real.angle) : - o.oangle (o.rotation θ x) x = -θ := -by simp [hx] - -/-- A vector has an angle of `θ` from the rotation of that vector by `θ`. -/ -@[simp] lemma oangle_rotation_self_right {x : V} (hx : x ≠ 0) (θ : real.angle) : - o.oangle x (o.rotation θ x) = θ := -by simp [hx] - -/-- Rotating the first vector by the angle between the two vectors results an an angle of 0. -/ -@[simp] lemma oangle_rotation_oangle_left (x y : V) : - o.oangle (o.rotation (o.oangle x y) x) y = 0 := -begin - by_cases hx : x = 0, - { simp [hx] }, - { by_cases hy : y = 0, - { simp [hy] }, - { simp [hx, hy] } } -end - -/-- Rotating the first vector by the angle between the two vectors and swapping the vectors -results an an angle of 0. -/ -@[simp] lemma oangle_rotation_oangle_right (x y : V) : - o.oangle y (o.rotation (o.oangle x y) x) = 0 := -begin - rw [oangle_rev], - simp -end - -/-- Rotating both vectors by the same angle does not change the angle between those vectors. -/ -@[simp] lemma oangle_rotation (x y : V) (θ : real.angle) : - o.oangle (o.rotation θ x) (o.rotation θ y) = o.oangle x y := -begin - by_cases hx : x = 0; by_cases hy : y = 0; - simp [hx, hy] -end - -/-- A rotation of a nonzero vector equals that vector if and only if the angle is zero. -/ -@[simp] lemma rotation_eq_self_iff_angle_eq_zero {x : V} (hx : x ≠ 0) (θ : real.angle) : - o.rotation θ x = x ↔ θ = 0 := -begin - split, - { intro h, - rw eq_comm, - simpa [hx, h] using o.oangle_rotation_right hx hx θ }, - { intro h, - simp [h] } -end - -/-- A nonzero vector equals a rotation of that vector if and only if the angle is zero. -/ -@[simp] lemma eq_rotation_self_iff_angle_eq_zero {x : V} (hx : x ≠ 0) (θ : real.angle) : - x = o.rotation θ x ↔ θ = 0 := -by rw [←o.rotation_eq_self_iff_angle_eq_zero hx, eq_comm] - -/-- A rotation of a vector equals that vector if and only if the vector or the angle is zero. -/ -lemma rotation_eq_self_iff (x : V) (θ : real.angle) : - o.rotation θ x = x ↔ x = 0 ∨ θ = 0 := -begin - by_cases h : x = 0; - simp [h] -end - -/-- A vector equals a rotation of that vector if and only if the vector or the angle is zero. -/ -lemma eq_rotation_self_iff (x : V) (θ : real.angle) : - x = o.rotation θ x ↔ x = 0 ∨ θ = 0 := -by rw [←rotation_eq_self_iff, eq_comm] - -/-- Rotating a vector by the angle to another vector gives the second vector if and only if the -norms are equal. -/ -@[simp] lemma rotation_oangle_eq_iff_norm_eq (x y : V) : - o.rotation (o.oangle x y) x = y ↔ ‖x‖ = ‖y‖ := -begin - split, - { intro h, - rw [←h, linear_isometry_equiv.norm_map] }, - { intro h, - rw o.eq_iff_oangle_eq_zero_of_norm_eq; - simp [h] } -end - -/-- The angle between two nonzero vectors is `θ` if and only if the second vector is the first -rotated by `θ` and scaled by the ratio of the norms. -/ -lemma oangle_eq_iff_eq_norm_div_norm_smul_rotation_of_ne_zero {x y : V} (hx : x ≠ 0) (hy : y ≠ 0) - (θ : real.angle) : o.oangle x y = θ ↔ y = (‖y‖ / ‖x‖) • o.rotation θ x := -begin - have hp := div_pos (norm_pos_iff.2 hy) (norm_pos_iff.2 hx), - split, - { rintro rfl, - rw [←linear_isometry_equiv.map_smul, ←o.oangle_smul_left_of_pos x y hp, - eq_comm, rotation_oangle_eq_iff_norm_eq, norm_smul, real.norm_of_nonneg hp.le, - div_mul_cancel _ (norm_ne_zero_iff.2 hx)] }, - { intro hye, - rw [hye, o.oangle_smul_right_of_pos _ _ hp, o.oangle_rotation_self_right hx] } -end - -/-- The angle between two nonzero vectors is `θ` if and only if the second vector is the first -rotated by `θ` and scaled by a positive real. -/ -lemma oangle_eq_iff_eq_pos_smul_rotation_of_ne_zero {x y : V} (hx : x ≠ 0) (hy : y ≠ 0) - (θ : real.angle) : o.oangle x y = θ ↔ ∃ r : ℝ, 0 < r ∧ y = r • o.rotation θ x := -begin - split, - { intro h, - rw o.oangle_eq_iff_eq_norm_div_norm_smul_rotation_of_ne_zero hx hy at h, - exact ⟨‖y‖ / ‖x‖, div_pos (norm_pos_iff.2 hy) (norm_pos_iff.2 hx), h⟩ }, - { rintro ⟨r, hr, rfl⟩, - rw [o.oangle_smul_right_of_pos _ _ hr, o.oangle_rotation_self_right hx] } -end - -/-- The angle between two vectors is `θ` if and only if they are nonzero and the second vector -is the first rotated by `θ` and scaled by the ratio of the norms, or `θ` and at least one of the -vectors are zero. -/ -lemma oangle_eq_iff_eq_norm_div_norm_smul_rotation_or_eq_zero {x y : V} (θ : real.angle) : - o.oangle x y = θ ↔ - (x ≠ 0 ∧ y ≠ 0 ∧ y = (‖y‖ / ‖x‖) • o.rotation θ x) ∨ (θ = 0 ∧ (x = 0 ∨ y = 0)) := -begin - by_cases hx : x = 0, - { simp [hx, eq_comm] }, - { by_cases hy : y = 0, - { simp [hy, eq_comm] }, - { rw o.oangle_eq_iff_eq_norm_div_norm_smul_rotation_of_ne_zero hx hy, - simp [hx, hy] } } -end - -/-- The angle between two vectors is `θ` if and only if they are nonzero and the second vector -is the first rotated by `θ` and scaled by a positive real, or `θ` and at least one of the -vectors are zero. -/ -lemma oangle_eq_iff_eq_pos_smul_rotation_or_eq_zero {x y : V} (θ : real.angle) : - o.oangle x y = θ ↔ - (x ≠ 0 ∧ y ≠ 0 ∧ ∃ r : ℝ, 0 < r ∧ y = r • o.rotation θ x) ∨ (θ = 0 ∧ (x = 0 ∨ y = 0)) := -begin - by_cases hx : x = 0, - { simp [hx, eq_comm] }, - { by_cases hy : y = 0, - { simp [hy, eq_comm] }, - { rw o.oangle_eq_iff_eq_pos_smul_rotation_of_ne_zero hx hy, - simp [hx, hy] } } -end - -/-- Any linear isometric equivalence in `V` with positive determinant is `rotation`. -/ -lemma exists_linear_isometry_equiv_eq_of_det_pos {f : V ≃ₗᵢ[ℝ] V} - (hd : 0 < (f.to_linear_equiv : V →ₗ[ℝ] V).det) : ∃ θ : real.angle, f = o.rotation θ := -begin - haveI : nontrivial V := - finite_dimensional.nontrivial_of_finrank_eq_succ (fact.out (finrank ℝ V = 2)), - obtain ⟨x, hx⟩ : ∃ x, x ≠ (0:V) := exists_ne (0:V), - use o.oangle x (f x), - apply linear_isometry_equiv.to_linear_equiv_injective, - apply linear_equiv.to_linear_map_injective, - apply (o.basis_right_angle_rotation x hx).ext, - intros i, - symmetry, - fin_cases i, - { simp }, - have : o.oangle (J x) (f (J x)) = o.oangle x (f x), - { simp only [oangle, o.linear_isometry_equiv_comp_right_angle_rotation f hd, - o.kahler_comp_right_angle_rotation] }, - simp [← this], -end - /-- The angle between two vectors, with respect to an orientation given by `orientation.map` with a linear isometric equivalence, equals the angle between those two vectors, transformed by the inverse of that equivalence, with respect to the original orientation. -/ @@ -929,34 +575,6 @@ end lemma oangle_neg_orientation_eq_neg (x y : V) : (-o).oangle x y = -(o.oangle x y) := by simp [oangle] -lemma rotation_map (θ : real.angle) (f : V ≃ₗᵢ[ℝ] V') (x : V') : - (orientation.map (fin 2) f.to_linear_equiv o).rotation θ x - = f (o.rotation θ (f.symm x)) := -by simp [rotation_apply, o.right_angle_rotation_map] - -@[simp] protected lemma _root_.complex.rotation (θ : real.angle) (z : ℂ) : - complex.orientation.rotation θ z = θ.exp_map_circle * z := -begin - simp only [rotation_apply, complex.right_angle_rotation, real.angle.coe_exp_map_circle, - real_smul], - ring -end - -/-- Rotation in an oriented real inner product space of dimension 2 can be evaluated in terms of a -complex-number representation of the space. -/ -lemma rotation_map_complex (θ : real.angle) (f : V ≃ₗᵢ[ℝ] ℂ) - (hf : (orientation.map (fin 2) f.to_linear_equiv o) = complex.orientation) (x : V) : - f (o.rotation θ x) = θ.exp_map_circle * f x := -begin - rw [← complex.rotation, ← hf, o.rotation_map], - simp, -end - -/-- Negating the orientation negates the angle in `rotation`. -/ -lemma rotation_neg_orientation_eq_neg (θ : real.angle) : - (-o).rotation θ = o.rotation (-θ) := -linear_isometry_equiv.ext $ by simp [rotation_apply] - /-- The inner product of two vectors is the product of the norms and the cosine of the oriented angle between the vectors. -/ lemma inner_eq_norm_mul_norm_mul_cos_oangle (x y : V) : @@ -1168,81 +786,6 @@ lemma inner_rev_eq_zero_of_oangle_eq_neg_pi_div_two {x y : V} (h : o.oangle x y ⟪y, x⟫ = 0 := by rw [real_inner_comm, o.inner_eq_zero_of_oangle_eq_neg_pi_div_two h] -/-- The inner product between a `π / 2` rotation of a vector and that vector is zero. -/ -@[simp] lemma inner_rotation_pi_div_two_left (x : V) : ⟪o.rotation (π / 2 : ℝ) x, x⟫ = 0 := -by rw [rotation_pi_div_two, inner_right_angle_rotation_self] - -/-- The inner product between a vector and a `π / 2` rotation of that vector is zero. -/ -@[simp] lemma inner_rotation_pi_div_two_right (x : V) : ⟪x, o.rotation (π / 2 : ℝ) x⟫ = 0 := -by rw [real_inner_comm, inner_rotation_pi_div_two_left] - -/-- The inner product between a multiple of a `π / 2` rotation of a vector and that vector is -zero. -/ -@[simp] lemma inner_smul_rotation_pi_div_two_left (x : V) (r : ℝ) : - ⟪r • o.rotation (π / 2 : ℝ) x, x⟫ = 0 := -by rw [inner_smul_left, inner_rotation_pi_div_two_left, mul_zero] - -/-- The inner product between a vector and a multiple of a `π / 2` rotation of that vector is -zero. -/ -@[simp] lemma inner_smul_rotation_pi_div_two_right (x : V) (r : ℝ) : - ⟪x, r • o.rotation (π / 2 : ℝ) x⟫ = 0 := -by rw [real_inner_comm, inner_smul_rotation_pi_div_two_left] - -/-- The inner product between a `π / 2` rotation of a vector and a multiple of that vector is -zero. -/ -@[simp] lemma inner_rotation_pi_div_two_left_smul (x : V) (r : ℝ) : - ⟪o.rotation (π / 2 : ℝ) x, r • x⟫ = 0 := -by rw [inner_smul_right, inner_rotation_pi_div_two_left, mul_zero] - -/-- The inner product between a multiple of a vector and a `π / 2` rotation of that vector is -zero. -/ -@[simp] lemma inner_rotation_pi_div_two_right_smul (x : V) (r : ℝ) : - ⟪r • x, o.rotation (π / 2 : ℝ) x⟫ = 0 := -by rw [real_inner_comm, inner_rotation_pi_div_two_left_smul] - -/-- The inner product between a multiple of a `π / 2` rotation of a vector and a multiple of -that vector is zero. -/ -@[simp] lemma inner_smul_rotation_pi_div_two_smul_left (x : V) (r₁ r₂ : ℝ) : - ⟪r₁ • o.rotation (π / 2 : ℝ) x, r₂ • x⟫ = 0 := -by rw [inner_smul_right, inner_smul_rotation_pi_div_two_left, mul_zero] - -/-- The inner product between a multiple of a vector and a multiple of a `π / 2` rotation of -that vector is zero. -/ -@[simp] lemma inner_smul_rotation_pi_div_two_smul_right (x : V) (r₁ r₂ : ℝ) : - ⟪r₂ • x, r₁ • o.rotation (π / 2 : ℝ) x⟫ = 0 := -by rw [real_inner_comm, inner_smul_rotation_pi_div_two_smul_left] - -/-- The inner product between two vectors is zero if and only if the first vector is zero or -the second is a multiple of a `π / 2` rotation of that vector. -/ -lemma inner_eq_zero_iff_eq_zero_or_eq_smul_rotation_pi_div_two {x y : V} : - ⟪x, y⟫ = 0 ↔ (x = 0 ∨ ∃ r : ℝ, r • o.rotation (π / 2 : ℝ) x = y) := -begin - rw ←o.eq_zero_or_oangle_eq_iff_inner_eq_zero, - refine ⟨λ h, _, λ h, _⟩, - { rcases h with rfl | rfl | h | h, - { exact or.inl rfl }, - { exact or.inr ⟨0, zero_smul _ _⟩ }, - { obtain ⟨r, hr, rfl⟩ := (o.oangle_eq_iff_eq_pos_smul_rotation_of_ne_zero - (o.left_ne_zero_of_oangle_eq_pi_div_two h) - (o.right_ne_zero_of_oangle_eq_pi_div_two h) _).1 h, - exact or.inr ⟨r, rfl⟩ }, - { obtain ⟨r, hr, rfl⟩ := (o.oangle_eq_iff_eq_pos_smul_rotation_of_ne_zero - (o.left_ne_zero_of_oangle_eq_neg_pi_div_two h) - (o.right_ne_zero_of_oangle_eq_neg_pi_div_two h) _).1 h, - refine or.inr ⟨-r, _⟩, - rw [neg_smul, ←smul_neg, o.neg_rotation_pi_div_two] } }, - { rcases h with rfl | ⟨r, rfl⟩, - { exact or.inl rfl }, - { by_cases hx : x = 0, { exact or.inl hx }, - rcases lt_trichotomy r 0 with hr | rfl | hr, - { refine or.inr (or.inr (or.inr _)), - rw [o.oangle_smul_right_of_neg _ _ hr, o.neg_rotation_pi_div_two, - o.oangle_rotation_self_right hx] }, - { exact or.inr (or.inl (zero_smul _ _)) }, - { refine or.inr (or.inr (or.inl _)), - rw [o.oangle_smul_right_of_pos _ _ hr, o.oangle_rotation_self_right hx] } } } -end - /-- Negating the first vector passed to `oangle` negates the sign of the angle. -/ @[simp] lemma oangle_sign_neg_left (x y : V) : (o.oangle (-x) y).sign = -((o.oangle x y).sign) := diff --git a/src/geometry/euclidean/angle/oriented/rotation.lean b/src/geometry/euclidean/angle/oriented/rotation.lean new file mode 100644 index 0000000000000..588502241a5e2 --- /dev/null +++ b/src/geometry/euclidean/angle/oriented/rotation.lean @@ -0,0 +1,488 @@ +/- +Copyright (c) 2022 Joseph Myers. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joseph Myers, Heather Macbeth +-/ +import geometry.euclidean.angle.oriented.basic + +/-! +# Rotations by oriented angles. + +This file defines rotations by oriented angles in real inner product spaces. + +## Main definitions + +* `orientation.rotation` is the rotation by an oriented angle with respect to an orientation. + +-/ + +noncomputable theory + +open finite_dimensional complex +open_locale real real_inner_product_space complex_conjugate + +namespace orientation + +local attribute [instance] fact_finite_dimensional_of_finrank_eq_succ +local attribute [instance] complex.finrank_real_complex_fact + +variables {V V' : Type*} [inner_product_space ℝ V] [inner_product_space ℝ V'] +variables [fact (finrank ℝ V = 2)] [fact (finrank ℝ V' = 2)] (o : orientation ℝ V (fin 2)) + +local notation `J` := o.right_angle_rotation + +/-- Auxiliary construction to build a rotation by the oriented angle `θ`. -/ +def rotation_aux (θ : real.angle) : V →ₗᵢ[ℝ] V := +linear_map.isometry_of_inner + (real.angle.cos θ • linear_map.id + + real.angle.sin θ • ↑(linear_isometry_equiv.to_linear_equiv J)) + begin + intros x y, + simp only [is_R_or_C.conj_to_real, id.def, linear_map.smul_apply, linear_map.add_apply, + linear_map.id_coe, linear_equiv.coe_coe, linear_isometry_equiv.coe_to_linear_equiv, + orientation.area_form_right_angle_rotation_left, + orientation.inner_right_angle_rotation_left, + orientation.inner_right_angle_rotation_right, + inner_add_left, inner_smul_left, inner_add_right, inner_smul_right], + linear_combination inner x y * θ.cos_sq_add_sin_sq, + end + +@[simp] lemma rotation_aux_apply (θ : real.angle) (x : V) : + o.rotation_aux θ x = real.angle.cos θ • x + real.angle.sin θ • J x := +rfl + +/-- A rotation by the oriented angle `θ`. -/ +def rotation (θ : real.angle) : V ≃ₗᵢ[ℝ] V := +linear_isometry_equiv.of_linear_isometry + (o.rotation_aux θ) + (real.angle.cos θ • linear_map.id - real.angle.sin θ • ↑(linear_isometry_equiv.to_linear_equiv J)) + begin + ext x, + convert congr_arg (λ t : ℝ, t • x) θ.cos_sq_add_sin_sq using 1, + { simp only [o.right_angle_rotation_right_angle_rotation, o.rotation_aux_apply, + function.comp_app, id.def, linear_equiv.coe_coe, linear_isometry.coe_to_linear_map, + linear_isometry_equiv.coe_to_linear_equiv, map_smul, map_sub, linear_map.coe_comp, + linear_map.id_coe, linear_map.smul_apply, linear_map.sub_apply, ← mul_smul, add_smul, + smul_add, smul_neg, smul_sub, mul_comm, sq], + abel }, + { simp }, + end + begin + ext x, + convert congr_arg (λ t : ℝ, t • x) θ.cos_sq_add_sin_sq using 1, + { simp only [o.right_angle_rotation_right_angle_rotation, o.rotation_aux_apply, + function.comp_app, id.def, linear_equiv.coe_coe, linear_isometry.coe_to_linear_map, + linear_isometry_equiv.coe_to_linear_equiv, map_add, map_smul, linear_map.coe_comp, + linear_map.id_coe, linear_map.smul_apply, linear_map.sub_apply, add_smul, ← mul_smul, + mul_comm, smul_add, smul_neg, sq], + abel }, + { simp }, + end + +lemma rotation_apply (θ : real.angle) (x : V) : + o.rotation θ x = real.angle.cos θ • x + real.angle.sin θ • J x := +rfl + +lemma rotation_symm_apply (θ : real.angle) (x : V) : + (o.rotation θ).symm x = real.angle.cos θ • x - real.angle.sin θ • J x := +rfl + +attribute [irreducible] rotation + +lemma rotation_eq_matrix_to_lin (θ : real.angle) {x : V} (hx : x ≠ 0) : + (o.rotation θ).to_linear_map + = matrix.to_lin + (o.basis_right_angle_rotation x hx) (o.basis_right_angle_rotation x hx) + !![θ.cos, -θ.sin; θ.sin, θ.cos] := +begin + apply (o.basis_right_angle_rotation x hx).ext, + intros i, + fin_cases i, + { rw matrix.to_lin_self, + simp [rotation_apply, fin.sum_univ_succ] }, + { rw matrix.to_lin_self, + simp [rotation_apply, fin.sum_univ_succ, add_comm] }, +end + +/-- The determinant of `rotation` (as a linear map) is equal to `1`. -/ +@[simp] lemma det_rotation (θ : real.angle) : + (o.rotation θ).to_linear_map.det = 1 := +begin + haveI : nontrivial V := + finite_dimensional.nontrivial_of_finrank_eq_succ (fact.out (finrank ℝ V = 2)), + obtain ⟨x, hx⟩ : ∃ x, x ≠ (0:V) := exists_ne (0:V), + rw o.rotation_eq_matrix_to_lin θ hx, + simpa [sq] using θ.cos_sq_add_sin_sq, +end + +/-- The determinant of `rotation` (as a linear equiv) is equal to `1`. -/ +@[simp] lemma linear_equiv_det_rotation (θ : real.angle) : + (o.rotation θ).to_linear_equiv.det = 1 := +units.ext $ o.det_rotation θ + +/-- The inverse of `rotation` is rotation by the negation of the angle. -/ +@[simp] lemma rotation_symm (θ : real.angle) : (o.rotation θ).symm = o.rotation (-θ) := +by ext; simp [o.rotation_apply, o.rotation_symm_apply, sub_eq_add_neg] + +/-- Rotation by 0 is the identity. -/ +@[simp] lemma rotation_zero : o.rotation 0 = linear_isometry_equiv.refl ℝ V := +by ext; simp [rotation] + +/-- Rotation by π is negation. -/ +@[simp] lemma rotation_pi : o.rotation π = linear_isometry_equiv.neg ℝ := +begin + ext x, + simp [rotation] +end + +/-- Rotation by π is negation. -/ +lemma rotation_pi_apply (x : V) : o.rotation π x = -x := +by simp + +/-- Rotation by π / 2 is the "right-angle-rotation" map `J`. -/ +lemma rotation_pi_div_two : o.rotation (π / 2 : ℝ) = J := +begin + ext x, + simp [rotation], +end + +/-- Rotating twice is equivalent to rotating by the sum of the angles. -/ +@[simp] lemma rotation_rotation (θ₁ θ₂ : real.angle) (x : V) : + o.rotation θ₁ (o.rotation θ₂ x) = o.rotation (θ₁ + θ₂) x := +begin + simp only [o.rotation_apply, ←mul_smul, real.angle.cos_add, real.angle.sin_add, add_smul, + sub_smul, linear_isometry_equiv.trans_apply, smul_add, linear_isometry_equiv.map_add, + linear_isometry_equiv.map_smul, right_angle_rotation_right_angle_rotation, smul_neg], + ring_nf, + abel, +end + +/-- Rotating twice is equivalent to rotating by the sum of the angles. -/ +@[simp] lemma rotation_trans (θ₁ θ₂ : real.angle) : + (o.rotation θ₁).trans (o.rotation θ₂) = o.rotation (θ₂ + θ₁) := +linear_isometry_equiv.ext $ λ _, by rw [←rotation_rotation, linear_isometry_equiv.trans_apply] + +/-- Rotating the first of two vectors by `θ` scales their Kahler form by `cos θ - sin θ * I`. -/ +@[simp] lemma kahler_rotation_left (x y : V) (θ : real.angle) : + o.kahler (o.rotation θ x) y = conj (θ.exp_map_circle : ℂ) * o.kahler x y := +begin + simp only [o.rotation_apply, map_add, map_mul, linear_map.map_smulₛₗ, ring_hom.id_apply, + linear_map.add_apply, linear_map.smul_apply, real_smul, kahler_right_angle_rotation_left, + real.angle.coe_exp_map_circle, is_R_or_C.conj_of_real, conj_I], + ring, +end + +/-- Negating a rotation is equivalent to rotation by π plus the angle. -/ +lemma neg_rotation (θ : real.angle) (x : V) : -o.rotation θ x = o.rotation (π + θ) x := +by rw [←o.rotation_pi_apply, rotation_rotation] + +/-- Negating a rotation by -π / 2 is equivalent to rotation by π / 2. -/ +@[simp] lemma neg_rotation_neg_pi_div_two (x : V) : + -o.rotation (-π / 2 : ℝ) x = o.rotation (π / 2 : ℝ) x := +by rw [neg_rotation, ←real.angle.coe_add, neg_div, ←sub_eq_add_neg, sub_half] + +/-- Negating a rotation by π / 2 is equivalent to rotation by -π / 2. -/ +lemma neg_rotation_pi_div_two (x : V) : -o.rotation (π / 2 : ℝ) x = o.rotation (-π / 2 : ℝ) x := +neg_eq_iff_neg_eq.1 $ o.neg_rotation_neg_pi_div_two _ + +/-- Rotating the first of two vectors by `θ` scales their Kahler form by `cos (-θ) + sin (-θ) * I`. +-/ +lemma kahler_rotation_left' (x y : V) (θ : real.angle) : + o.kahler (o.rotation θ x) y = (-θ).exp_map_circle * o.kahler x y := +by simpa [coe_inv_circle_eq_conj, -kahler_rotation_left] using o.kahler_rotation_left x y θ + +/-- Rotating the second of two vectors by `θ` scales their Kahler form by `cos θ + sin θ * I`. -/ +@[simp] lemma kahler_rotation_right (x y : V) (θ : real.angle) : + o.kahler x (o.rotation θ y) = θ.exp_map_circle * o.kahler x y := +begin + simp only [o.rotation_apply, map_add, linear_map.map_smulₛₗ, ring_hom.id_apply, real_smul, + kahler_right_angle_rotation_right, real.angle.coe_exp_map_circle], + ring, +end + +/-- Rotating the first vector by `θ` subtracts `θ` from the angle between two vectors. -/ +@[simp] lemma oangle_rotation_left {x y : V} (hx : x ≠ 0) (hy : y ≠ 0) (θ : real.angle) : + o.oangle (o.rotation θ x) y = o.oangle x y - θ := +begin + simp only [oangle, o.kahler_rotation_left'], + rw [complex.arg_mul_coe_angle, real.angle.arg_exp_map_circle], + { abel }, + { exact ne_zero_of_mem_circle _ }, + { exact o.kahler_ne_zero hx hy }, +end + +/-- Rotating the second vector by `θ` adds `θ` to the angle between two vectors. -/ +@[simp] lemma oangle_rotation_right {x y : V} (hx : x ≠ 0) (hy : y ≠ 0) (θ : real.angle) : + o.oangle x (o.rotation θ y) = o.oangle x y + θ := +begin + simp only [oangle, o.kahler_rotation_right], + rw [complex.arg_mul_coe_angle, real.angle.arg_exp_map_circle], + { abel }, + { exact ne_zero_of_mem_circle _ }, + { exact o.kahler_ne_zero hx hy }, +end + +/-- The rotation of a vector by `θ` has an angle of `-θ` from that vector. -/ +@[simp] lemma oangle_rotation_self_left {x : V} (hx : x ≠ 0) (θ : real.angle) : + o.oangle (o.rotation θ x) x = -θ := +by simp [hx] + +/-- A vector has an angle of `θ` from the rotation of that vector by `θ`. -/ +@[simp] lemma oangle_rotation_self_right {x : V} (hx : x ≠ 0) (θ : real.angle) : + o.oangle x (o.rotation θ x) = θ := +by simp [hx] + +/-- Rotating the first vector by the angle between the two vectors results an an angle of 0. -/ +@[simp] lemma oangle_rotation_oangle_left (x y : V) : + o.oangle (o.rotation (o.oangle x y) x) y = 0 := +begin + by_cases hx : x = 0, + { simp [hx] }, + { by_cases hy : y = 0, + { simp [hy] }, + { simp [hx, hy] } } +end + +/-- Rotating the first vector by the angle between the two vectors and swapping the vectors +results an an angle of 0. -/ +@[simp] lemma oangle_rotation_oangle_right (x y : V) : + o.oangle y (o.rotation (o.oangle x y) x) = 0 := +begin + rw [oangle_rev], + simp +end + +/-- Rotating both vectors by the same angle does not change the angle between those vectors. -/ +@[simp] lemma oangle_rotation (x y : V) (θ : real.angle) : + o.oangle (o.rotation θ x) (o.rotation θ y) = o.oangle x y := +begin + by_cases hx : x = 0; by_cases hy : y = 0; + simp [hx, hy] +end + +/-- A rotation of a nonzero vector equals that vector if and only if the angle is zero. -/ +@[simp] lemma rotation_eq_self_iff_angle_eq_zero {x : V} (hx : x ≠ 0) (θ : real.angle) : + o.rotation θ x = x ↔ θ = 0 := +begin + split, + { intro h, + rw eq_comm, + simpa [hx, h] using o.oangle_rotation_right hx hx θ }, + { intro h, + simp [h] } +end + +/-- A nonzero vector equals a rotation of that vector if and only if the angle is zero. -/ +@[simp] lemma eq_rotation_self_iff_angle_eq_zero {x : V} (hx : x ≠ 0) (θ : real.angle) : + x = o.rotation θ x ↔ θ = 0 := +by rw [←o.rotation_eq_self_iff_angle_eq_zero hx, eq_comm] + +/-- A rotation of a vector equals that vector if and only if the vector or the angle is zero. -/ +lemma rotation_eq_self_iff (x : V) (θ : real.angle) : + o.rotation θ x = x ↔ x = 0 ∨ θ = 0 := +begin + by_cases h : x = 0; + simp [h] +end + +/-- A vector equals a rotation of that vector if and only if the vector or the angle is zero. -/ +lemma eq_rotation_self_iff (x : V) (θ : real.angle) : + x = o.rotation θ x ↔ x = 0 ∨ θ = 0 := +by rw [←rotation_eq_self_iff, eq_comm] + +/-- Rotating a vector by the angle to another vector gives the second vector if and only if the +norms are equal. -/ +@[simp] lemma rotation_oangle_eq_iff_norm_eq (x y : V) : + o.rotation (o.oangle x y) x = y ↔ ‖x‖ = ‖y‖ := +begin + split, + { intro h, + rw [←h, linear_isometry_equiv.norm_map] }, + { intro h, + rw o.eq_iff_oangle_eq_zero_of_norm_eq; + simp [h] } +end + +/-- The angle between two nonzero vectors is `θ` if and only if the second vector is the first +rotated by `θ` and scaled by the ratio of the norms. -/ +lemma oangle_eq_iff_eq_norm_div_norm_smul_rotation_of_ne_zero {x y : V} (hx : x ≠ 0) (hy : y ≠ 0) + (θ : real.angle) : o.oangle x y = θ ↔ y = (‖y‖ / ‖x‖) • o.rotation θ x := +begin + have hp := div_pos (norm_pos_iff.2 hy) (norm_pos_iff.2 hx), + split, + { rintro rfl, + rw [←linear_isometry_equiv.map_smul, ←o.oangle_smul_left_of_pos x y hp, + eq_comm, rotation_oangle_eq_iff_norm_eq, norm_smul, real.norm_of_nonneg hp.le, + div_mul_cancel _ (norm_ne_zero_iff.2 hx)] }, + { intro hye, + rw [hye, o.oangle_smul_right_of_pos _ _ hp, o.oangle_rotation_self_right hx] } +end + +/-- The angle between two nonzero vectors is `θ` if and only if the second vector is the first +rotated by `θ` and scaled by a positive real. -/ +lemma oangle_eq_iff_eq_pos_smul_rotation_of_ne_zero {x y : V} (hx : x ≠ 0) (hy : y ≠ 0) + (θ : real.angle) : o.oangle x y = θ ↔ ∃ r : ℝ, 0 < r ∧ y = r • o.rotation θ x := +begin + split, + { intro h, + rw o.oangle_eq_iff_eq_norm_div_norm_smul_rotation_of_ne_zero hx hy at h, + exact ⟨‖y‖ / ‖x‖, div_pos (norm_pos_iff.2 hy) (norm_pos_iff.2 hx), h⟩ }, + { rintro ⟨r, hr, rfl⟩, + rw [o.oangle_smul_right_of_pos _ _ hr, o.oangle_rotation_self_right hx] } +end + +/-- The angle between two vectors is `θ` if and only if they are nonzero and the second vector +is the first rotated by `θ` and scaled by the ratio of the norms, or `θ` and at least one of the +vectors are zero. -/ +lemma oangle_eq_iff_eq_norm_div_norm_smul_rotation_or_eq_zero {x y : V} (θ : real.angle) : + o.oangle x y = θ ↔ + (x ≠ 0 ∧ y ≠ 0 ∧ y = (‖y‖ / ‖x‖) • o.rotation θ x) ∨ (θ = 0 ∧ (x = 0 ∨ y = 0)) := +begin + by_cases hx : x = 0, + { simp [hx, eq_comm] }, + { by_cases hy : y = 0, + { simp [hy, eq_comm] }, + { rw o.oangle_eq_iff_eq_norm_div_norm_smul_rotation_of_ne_zero hx hy, + simp [hx, hy] } } +end + +/-- The angle between two vectors is `θ` if and only if they are nonzero and the second vector +is the first rotated by `θ` and scaled by a positive real, or `θ` and at least one of the +vectors are zero. -/ +lemma oangle_eq_iff_eq_pos_smul_rotation_or_eq_zero {x y : V} (θ : real.angle) : + o.oangle x y = θ ↔ + (x ≠ 0 ∧ y ≠ 0 ∧ ∃ r : ℝ, 0 < r ∧ y = r • o.rotation θ x) ∨ (θ = 0 ∧ (x = 0 ∨ y = 0)) := +begin + by_cases hx : x = 0, + { simp [hx, eq_comm] }, + { by_cases hy : y = 0, + { simp [hy, eq_comm] }, + { rw o.oangle_eq_iff_eq_pos_smul_rotation_of_ne_zero hx hy, + simp [hx, hy] } } +end + +/-- Any linear isometric equivalence in `V` with positive determinant is `rotation`. -/ +lemma exists_linear_isometry_equiv_eq_of_det_pos {f : V ≃ₗᵢ[ℝ] V} + (hd : 0 < (f.to_linear_equiv : V →ₗ[ℝ] V).det) : ∃ θ : real.angle, f = o.rotation θ := +begin + haveI : nontrivial V := + finite_dimensional.nontrivial_of_finrank_eq_succ (fact.out (finrank ℝ V = 2)), + obtain ⟨x, hx⟩ : ∃ x, x ≠ (0:V) := exists_ne (0:V), + use o.oangle x (f x), + apply linear_isometry_equiv.to_linear_equiv_injective, + apply linear_equiv.to_linear_map_injective, + apply (o.basis_right_angle_rotation x hx).ext, + intros i, + symmetry, + fin_cases i, + { simp }, + have : o.oangle (J x) (f (J x)) = o.oangle x (f x), + { simp only [oangle, o.linear_isometry_equiv_comp_right_angle_rotation f hd, + o.kahler_comp_right_angle_rotation] }, + simp [← this], +end + +lemma rotation_map (θ : real.angle) (f : V ≃ₗᵢ[ℝ] V') (x : V') : + (orientation.map (fin 2) f.to_linear_equiv o).rotation θ x + = f (o.rotation θ (f.symm x)) := +by simp [rotation_apply, o.right_angle_rotation_map] + +@[simp] protected lemma _root_.complex.rotation (θ : real.angle) (z : ℂ) : + complex.orientation.rotation θ z = θ.exp_map_circle * z := +begin + simp only [rotation_apply, complex.right_angle_rotation, real.angle.coe_exp_map_circle, + real_smul], + ring +end + +/-- Rotation in an oriented real inner product space of dimension 2 can be evaluated in terms of a +complex-number representation of the space. -/ +lemma rotation_map_complex (θ : real.angle) (f : V ≃ₗᵢ[ℝ] ℂ) + (hf : (orientation.map (fin 2) f.to_linear_equiv o) = complex.orientation) (x : V) : + f (o.rotation θ x) = θ.exp_map_circle * f x := +begin + rw [← complex.rotation, ← hf, o.rotation_map], + simp, +end + +/-- Negating the orientation negates the angle in `rotation`. -/ +lemma rotation_neg_orientation_eq_neg (θ : real.angle) : + (-o).rotation θ = o.rotation (-θ) := +linear_isometry_equiv.ext $ by simp [rotation_apply] + +/-- The inner product between a `π / 2` rotation of a vector and that vector is zero. -/ +@[simp] lemma inner_rotation_pi_div_two_left (x : V) : ⟪o.rotation (π / 2 : ℝ) x, x⟫ = 0 := +by rw [rotation_pi_div_two, inner_right_angle_rotation_self] + +/-- The inner product between a vector and a `π / 2` rotation of that vector is zero. -/ +@[simp] lemma inner_rotation_pi_div_two_right (x : V) : ⟪x, o.rotation (π / 2 : ℝ) x⟫ = 0 := +by rw [real_inner_comm, inner_rotation_pi_div_two_left] + +/-- The inner product between a multiple of a `π / 2` rotation of a vector and that vector is +zero. -/ +@[simp] lemma inner_smul_rotation_pi_div_two_left (x : V) (r : ℝ) : + ⟪r • o.rotation (π / 2 : ℝ) x, x⟫ = 0 := +by rw [inner_smul_left, inner_rotation_pi_div_two_left, mul_zero] + +/-- The inner product between a vector and a multiple of a `π / 2` rotation of that vector is +zero. -/ +@[simp] lemma inner_smul_rotation_pi_div_two_right (x : V) (r : ℝ) : + ⟪x, r • o.rotation (π / 2 : ℝ) x⟫ = 0 := +by rw [real_inner_comm, inner_smul_rotation_pi_div_two_left] + +/-- The inner product between a `π / 2` rotation of a vector and a multiple of that vector is +zero. -/ +@[simp] lemma inner_rotation_pi_div_two_left_smul (x : V) (r : ℝ) : + ⟪o.rotation (π / 2 : ℝ) x, r • x⟫ = 0 := +by rw [inner_smul_right, inner_rotation_pi_div_two_left, mul_zero] + +/-- The inner product between a multiple of a vector and a `π / 2` rotation of that vector is +zero. -/ +@[simp] lemma inner_rotation_pi_div_two_right_smul (x : V) (r : ℝ) : + ⟪r • x, o.rotation (π / 2 : ℝ) x⟫ = 0 := +by rw [real_inner_comm, inner_rotation_pi_div_two_left_smul] + +/-- The inner product between a multiple of a `π / 2` rotation of a vector and a multiple of +that vector is zero. -/ +@[simp] lemma inner_smul_rotation_pi_div_two_smul_left (x : V) (r₁ r₂ : ℝ) : + ⟪r₁ • o.rotation (π / 2 : ℝ) x, r₂ • x⟫ = 0 := +by rw [inner_smul_right, inner_smul_rotation_pi_div_two_left, mul_zero] + +/-- The inner product between a multiple of a vector and a multiple of a `π / 2` rotation of +that vector is zero. -/ +@[simp] lemma inner_smul_rotation_pi_div_two_smul_right (x : V) (r₁ r₂ : ℝ) : + ⟪r₂ • x, r₁ • o.rotation (π / 2 : ℝ) x⟫ = 0 := +by rw [real_inner_comm, inner_smul_rotation_pi_div_two_smul_left] + +/-- The inner product between two vectors is zero if and only if the first vector is zero or +the second is a multiple of a `π / 2` rotation of that vector. -/ +lemma inner_eq_zero_iff_eq_zero_or_eq_smul_rotation_pi_div_two {x y : V} : + ⟪x, y⟫ = 0 ↔ (x = 0 ∨ ∃ r : ℝ, r • o.rotation (π / 2 : ℝ) x = y) := +begin + rw ←o.eq_zero_or_oangle_eq_iff_inner_eq_zero, + refine ⟨λ h, _, λ h, _⟩, + { rcases h with rfl | rfl | h | h, + { exact or.inl rfl }, + { exact or.inr ⟨0, zero_smul _ _⟩ }, + { obtain ⟨r, hr, rfl⟩ := (o.oangle_eq_iff_eq_pos_smul_rotation_of_ne_zero + (o.left_ne_zero_of_oangle_eq_pi_div_two h) + (o.right_ne_zero_of_oangle_eq_pi_div_two h) _).1 h, + exact or.inr ⟨r, rfl⟩ }, + { obtain ⟨r, hr, rfl⟩ := (o.oangle_eq_iff_eq_pos_smul_rotation_of_ne_zero + (o.left_ne_zero_of_oangle_eq_neg_pi_div_two h) + (o.right_ne_zero_of_oangle_eq_neg_pi_div_two h) _).1 h, + refine or.inr ⟨-r, _⟩, + rw [neg_smul, ←smul_neg, o.neg_rotation_pi_div_two] } }, + { rcases h with rfl | ⟨r, rfl⟩, + { exact or.inl rfl }, + { by_cases hx : x = 0, { exact or.inl hx }, + rcases lt_trichotomy r 0 with hr | rfl | hr, + { refine or.inr (or.inr (or.inr _)), + rw [o.oangle_smul_right_of_neg _ _ hr, o.neg_rotation_pi_div_two, + o.oangle_rotation_self_right hx] }, + { exact or.inr (or.inl (zero_smul _ _)) }, + { refine or.inr (or.inr (or.inl _)), + rw [o.oangle_smul_right_of_pos _ _ hr, o.oangle_rotation_self_right hx] } } } +end + +end orientation From 509de852e1de55e1efa8eacfa11df0823f26f226 Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Wed, 18 Jan 2023 17:47:29 +0000 Subject: [PATCH 0289/1291] fix(geometry/euclidean/circumcenter): fix `fintype_finite` lint error (#18213) Change `affine_independent.exists_unique_dist_eq` to use `finite` instead of `fintype`. --- src/geometry/euclidean/circumcenter.lean | 5 +++-- 1 file changed, 3 insertions(+), 2 deletions(-) diff --git a/src/geometry/euclidean/circumcenter.lean b/src/geometry/euclidean/circumcenter.lean index f03398d2a09fc..da7ffb6e7c325 100644 --- a/src/geometry/euclidean/circumcenter.lean +++ b/src/geometry/euclidean/circumcenter.lean @@ -188,10 +188,11 @@ end /-- Given a finite nonempty affinely independent family of points, there is a unique (circumcenter, circumradius) pair for those points in the affine subspace they span. -/ -lemma _root_.affine_independent.exists_unique_dist_eq {ι : Type*} [hne : nonempty ι] [fintype ι] +lemma _root_.affine_independent.exists_unique_dist_eq {ι : Type*} [hne : nonempty ι] [finite ι] {p : ι → P} (ha : affine_independent ℝ p) : ∃! cs : sphere P, cs.center ∈ affine_span ℝ (set.range p) ∧ set.range p ⊆ (cs : set P) := begin + casesI nonempty_fintype ι, unfreezingI { induction hn : fintype.card ι with m hm generalizing ι }, { exfalso, have h := fintype.card_pos_iff.2 hne, @@ -225,7 +226,7 @@ begin { simp } }, haveI : nonempty ι2 := fintype.card_pos_iff.1 (hc.symm ▸ nat.zero_lt_succ _), have ha2 : affine_independent ℝ (λ i2 : ι2, p i2) := ha.subtype _, - replace hm := hm ha2 hc, + replace hm := hm ha2 _ hc, have hr : set.range p = insert (p i) (set.range (λ i2 : ι2, p i2)), { change _ = insert _ (set.range (λ i2 : {x | x ≠ i}, p i2)), rw [←set.image_eq_range, ←set.image_univ, ←set.image_insert_eq], From e0e2f10d64d8a5fd11140de398eaa1322eb46c07 Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Wed, 18 Jan 2023 22:53:23 +0000 Subject: [PATCH 0290/1291] chore(analysis/normed_space/continuous_linear_map): fix docs typo (#18215) --- src/analysis/normed_space/continuous_linear_map.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/analysis/normed_space/continuous_linear_map.lean b/src/analysis/normed_space/continuous_linear_map.lean index bb3bc4ae41afb..50c0db7872b3f 100644 --- a/src/analysis/normed_space/continuous_linear_map.lean +++ b/src/analysis/normed_space/continuous_linear_map.lean @@ -175,7 +175,7 @@ linear_equiv.to_continuous_linear_equiv_of_bounds f a a⁻¹ end seminormed -/- ## The span of a single vector -/ +/-! ## The span of a single vector -/ section seminormed From c1ada9d195393af8a78376fb35d4d7108666bda9 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 19 Jan 2023 00:39:05 +0000 Subject: [PATCH 0291/1291] fix(tactic/ring): perform definitional rather than syntactic matches on `^` (#18156) This code is copied from a similar pattern for `has_div.div` a few lines up. Follows on from #18129 --- src/tactic/ring.lean | 25 +++++++++++++++---------- test/ring.lean | 7 +++++++ 2 files changed, 22 insertions(+), 10 deletions(-) diff --git a/src/tactic/ring.lean b/src/tactic/ring.lean index 59995a1c7e3d9..4dcb10f74db63 100644 --- a/src/tactic/ring.lean +++ b/src/tactic/ring.lean @@ -499,16 +499,21 @@ meta def eval (norm_atom : expr → tactic (expr × expr)) : expr → ring_m (ho p' ← ic_lift $ λ ic, ic.mk_app ``unfold_div [e₁, e₂, e', p], return (e', p')) (eval_norm_atom norm_atom e) -| e@`(@has_pow.pow _ _ %%P %%e₁ %%e₂) := do - (e₂', p₂) ← lift $ norm_num.derive e₂ <|> refl_conv e₂, - match e₂'.to_nat, P with - | some k, `(monoid.has_pow) := do - (e₁', p₁) ← eval e₁, - (e', p') ← eval_pow e₁' (e₂, k), - p ← ic_lift $ λ ic, ic.mk_app ``subst_into_pow [e₁, e₂, e₁', e₂', e', p₁, p₂, p'], - return (e', p) - | _, _ := eval_norm_atom norm_atom e - end +| e@`(@has_pow.pow _ _ %%inst %%e₁ %%e₂) := mcond + (succeeds (do + inst' ← ic_lift $ λ ic, ic.mk_app ``monoid.has_pow [], + lift $ is_def_eq inst inst')) + (do + (e₂', p₂) ← lift $ norm_num.derive e₂ <|> refl_conv e₂, + match e₂'.to_nat with + | some k := do + (e₁', p₁) ← eval e₁, + (e', p') ← eval_pow e₁' (e₂, k), + p ← ic_lift $ λ ic, ic.mk_app ``subst_into_pow [e₁, e₂, e₁', e₂', e', p₁, p₂, p'], + return (e', p) + | _ := eval_norm_atom norm_atom e + end) + (eval_norm_atom norm_atom e) | e := match e.to_nat with | some n := (const e n).refl_conv | none := eval_norm_atom norm_atom e diff --git a/test/ring.lean b/test/ring.lean index 82c99e1c65feb..f0d15a8954313 100644 --- a/test/ring.lean +++ b/test/ring.lean @@ -95,3 +95,10 @@ begin success_if_fail {{ ring_nf {recursive := ff} }}, ring_nf end + +-- instances do not have to syntactically be `monoid.has_pow` +example {R} [comm_semiring R] (x : ℕ → R) : x ^ 2 = x * x := by ring + +-- even if there's an instance we don't recognize, we treat it as an atom +example {R} [field R] (x : ℕ → R) : + (x ^ (2 : ℤ)) ^ 2 = (x ^ (2 : ℤ)) * (x ^ (2 : ℤ)) := by ring From 3b9e2ef291f497fde2951fd2dd208eb0a2c952af Mon Sep 17 00:00:00 2001 From: JovanGerb Date: Thu, 19 Jan 2023 03:14:13 +0000 Subject: [PATCH 0292/1291] feat(data/real/sqrt): add sqrt_div' lemma (#18216) The current sqrt_div lemma requires the numerator to be non-negative. It is also sufficient to have the denominator be non-negative. This matches how we have both `sqrt_mul` and `sqrt_mul'`. --- src/data/real/sqrt.lean | 3 +++ 1 file changed, 3 insertions(+) diff --git a/src/data/real/sqrt.lean b/src/data/real/sqrt.lean index 2ba270b9a1395..d2cfa45030bf5 100644 --- a/src/data/real/sqrt.lean +++ b/src/data/real/sqrt.lean @@ -290,6 +290,9 @@ by rw [sqrt, real.to_nnreal_inv, nnreal.sqrt_inv, nnreal.coe_inv, sqrt] @[simp] theorem sqrt_div (hx : 0 ≤ x) (y : ℝ) : sqrt (x / y) = sqrt x / sqrt y := by rw [division_def, sqrt_mul hx, sqrt_inv, division_def] +@[simp] theorem sqrt_div' (x) {y : ℝ} (hy : 0 ≤ y) : sqrt (x / y) = sqrt x / sqrt y := +by rw [division_def, sqrt_mul' x (inv_nonneg.2 hy), sqrt_inv, division_def] + @[simp] theorem div_sqrt : x / sqrt x = sqrt x := begin cases le_or_lt x 0, From 1684fd2a39e7a9aa1ba6a2d33ea2751d79fc37f2 Mon Sep 17 00:00:00 2001 From: leanprover-community-bot Date: Thu, 19 Jan 2023 05:03:33 +0000 Subject: [PATCH 0293/1291] chore(scripts): update nolints.txt (#18223) I am happy to remove some nolints for you! --- scripts/nolints.txt | 3 --- 1 file changed, 3 deletions(-) diff --git a/scripts/nolints.txt b/scripts/nolints.txt index 60d36e81c535d..0da58c64c895a 100644 --- a/scripts/nolints.txt +++ b/scripts/nolints.txt @@ -345,9 +345,6 @@ apply_nolint stream.unfolds doc_blame -- data/stream/init.lean apply_nolint stream.is_bisimulation doc_blame --- geometry/euclidean/circumcenter.lean -apply_nolint affine_independent.exists_unique_dist_eq fintype_finite - -- group_theory/group_action/sub_mul_action.lean apply_nolint sub_mul_action.has_zero fails_quickly From 3c4225288b55380a90df078ebae0991080b12393 Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Thu, 19 Jan 2023 07:18:17 +0000 Subject: [PATCH 0294/1291] refactor(analysis/normed/group/hom): split (#18219) Half of this file is completely elementary, able to be proved directly from the definitions in `normed/group/hom/basic` after a few instances are added there. The other half consists of technical lemmas from LTE, never used elsewhere in mathlib, and requires more imports. Since this file is imported by many files (including `data/complex/is_R_or_C`, see #18217 for a discussion of what that file imports), I propose splitting off the LTE half. --- src/analysis/normed/field/basic.lean | 52 +------- src/analysis/normed/group/basic.lean | 56 ++++++++ .../normed/group/controlled_closure.lean | 121 ++++++++++++++++++ src/analysis/normed/group/hom.lean | 110 +--------------- src/analysis/normed/group/quotient.lean | 1 + src/analysis/special_functions/exp.lean | 3 +- 6 files changed, 186 insertions(+), 157 deletions(-) create mode 100644 src/analysis/normed/group/controlled_closure.lean diff --git a/src/analysis/normed/field/basic.lean b/src/analysis/normed/field/basic.lean index 8b14076599c1b..9850a7146b863 100644 --- a/src/analysis/normed/field/basic.lean +++ b/src/analysis/normed/field/basic.lean @@ -6,7 +6,6 @@ Authors: Patrick Massot, Johannes Hölzl import algebra.algebra.subalgebra.basic import analysis.normed.group.basic import topology.instances.ennreal -import topology.instances.rat /-! # Normed fields @@ -676,62 +675,21 @@ lemma normed_add_comm_group.tendsto_at_top' [nonempty α] [semilattice_sup α] [ (at_top_basis_Ioi.tendsto_iff metric.nhds_basis_ball).trans (by simp [dist_eq_norm]) instance : normed_comm_ring ℤ := -{ norm := λ n, ‖(n : ℝ)‖, - norm_mul := λ m n, le_of_eq $ by simp only [norm, int.cast_mul, abs_mul], - dist_eq := λ m n, by simp only [int.dist_eq, norm, int.cast_sub], - mul_comm := mul_comm } - -@[norm_cast] lemma int.norm_cast_real (m : ℤ) : ‖(m : ℝ)‖ = ‖m‖ := rfl - -lemma int.norm_eq_abs (n : ℤ) : ‖n‖ = |n| := rfl - -@[simp] lemma int.norm_coe_nat (n : ℕ) : ‖(n : ℤ)‖ = n := by simp [int.norm_eq_abs] - -lemma nnreal.coe_nat_abs (n : ℤ) : (n.nat_abs : ℝ≥0) = ‖n‖₊ := -nnreal.eq $ calc ((n.nat_abs : ℝ≥0) : ℝ) - = (n.nat_abs : ℤ) : by simp only [int.cast_coe_nat, nnreal.coe_nat_cast] - ... = |n| : by simp only [int.coe_nat_abs, int.cast_abs] - ... = ‖n‖ : rfl - -lemma int.abs_le_floor_nnreal_iff (z : ℤ) (c : ℝ≥0) : |z| ≤ ⌊c⌋₊ ↔ ‖z‖₊ ≤ c := -begin - rw [int.abs_eq_nat_abs, int.coe_nat_le, nat.le_floor_iff (zero_le c)], - congr', - exact nnreal.coe_nat_abs z, -end +{ norm_mul := λ m n, le_of_eq $ by simp only [norm, int.cast_mul, abs_mul], + mul_comm := mul_comm, + .. int.normed_add_comm_group } instance : norm_one_class ℤ := ⟨by simp [← int.norm_cast_real]⟩ instance : normed_field ℚ := -{ norm := λ r, ‖(r : ℝ)‖, - norm_mul' := λ r₁ r₂, by simp only [norm, rat.cast_mul, abs_mul], - dist_eq := λ r₁ r₂, by simp only [rat.dist_eq, norm, rat.cast_sub] } +{ norm_mul' := λ r₁ r₂, by simp only [norm, rat.cast_mul, abs_mul], + .. rat.normed_add_comm_group } instance : densely_normed_field ℚ := { lt_norm_lt := λ r₁ r₂ h₀ hr, let ⟨q, h⟩ := exists_rat_btwn hr in ⟨q, by { unfold norm, rwa abs_of_pos (h₀.trans_lt h.1) } ⟩ } -@[norm_cast, simp] lemma rat.norm_cast_real (r : ℚ) : ‖(r : ℝ)‖ = ‖r‖ := rfl - -@[norm_cast, simp] lemma int.norm_cast_rat (m : ℤ) : ‖(m : ℚ)‖ = ‖m‖ := -by rw [← rat.norm_cast_real, ← int.norm_cast_real]; congr' 1; norm_cast - --- Now that we've installed the norm on `ℤ`, --- we can state some lemmas about `zsmul`. -section -variables [seminormed_comm_group α] - -@[to_additive norm_zsmul_le] -lemma norm_zpow_le_mul_norm (n : ℤ) (a : α) : ‖a^n‖ ≤ ‖n‖ * ‖a‖ := -by rcases n.eq_coe_or_neg with ⟨n, rfl | rfl⟩; simpa using norm_pow_le_mul_norm n a - -@[to_additive nnnorm_zsmul_le] -lemma nnnorm_zpow_le_mul_norm (n : ℤ) (a : α) : ‖a^n‖₊ ≤ ‖n‖₊ * ‖a‖₊ := -by simpa only [← nnreal.coe_le_coe, nnreal.coe_mul] using norm_zpow_le_mul_norm n a - -end - section ring_hom_isometric variables {R₁ : Type*} {R₂ : Type*} {R₃ : Type*} diff --git a/src/analysis/normed/group/basic.lean b/src/analysis/normed/group/basic.lean index f82dcde4fcc56..32bdec02e5262 100644 --- a/src/analysis/normed/group/basic.lean +++ b/src/analysis/normed/group/basic.lean @@ -6,6 +6,7 @@ Authors: Patrick Massot, Johannes Hölzl, Yaël Dillies import analysis.normed.group.seminorm import order.liminf_limsup import topology.algebra.uniform_group +import topology.instances.rat import topology.metric_space.algebra import topology.metric_space.isometry import topology.sequences @@ -1271,6 +1272,61 @@ end end real +namespace int + +instance : normed_add_comm_group ℤ := +{ norm := λ n, ‖(n : ℝ)‖, + dist_eq := λ m n, by simp only [int.dist_eq, norm, int.cast_sub] } + +@[norm_cast] lemma norm_cast_real (m : ℤ) : ‖(m : ℝ)‖ = ‖m‖ := rfl + +lemma norm_eq_abs (n : ℤ) : ‖n‖ = |n| := rfl + +@[simp] lemma norm_coe_nat (n : ℕ) : ‖(n : ℤ)‖ = n := by simp [int.norm_eq_abs] + +lemma _root_.nnreal.coe_nat_abs (n : ℤ) : (n.nat_abs : ℝ≥0) = ‖n‖₊ := +nnreal.eq $ calc ((n.nat_abs : ℝ≥0) : ℝ) + = (n.nat_abs : ℤ) : by simp only [int.cast_coe_nat, nnreal.coe_nat_cast] + ... = |n| : by simp only [int.coe_nat_abs, int.cast_abs] + ... = ‖n‖ : rfl + +lemma abs_le_floor_nnreal_iff (z : ℤ) (c : ℝ≥0) : |z| ≤ ⌊c⌋₊ ↔ ‖z‖₊ ≤ c := +begin + rw [int.abs_eq_nat_abs, int.coe_nat_le, nat.le_floor_iff (zero_le c)], + congr', + exact nnreal.coe_nat_abs z, +end + +end int + +namespace rat + +instance : normed_add_comm_group ℚ := +{ norm := λ r, ‖(r : ℝ)‖, + dist_eq := λ r₁ r₂, by simp only [rat.dist_eq, norm, rat.cast_sub] } + +@[norm_cast, simp] lemma norm_cast_real (r : ℚ) : ‖(r : ℝ)‖ = ‖r‖ := rfl + +@[norm_cast, simp] lemma _root_.int.norm_cast_rat (m : ℤ) : ‖(m : ℚ)‖ = ‖m‖ := +by rw [← rat.norm_cast_real, ← int.norm_cast_real]; congr' 1; norm_cast + +end rat + +-- Now that we've installed the norm on `ℤ`, +-- we can state some lemmas about `zsmul`. +section +variables [seminormed_comm_group α] + +@[to_additive norm_zsmul_le] +lemma norm_zpow_le_mul_norm (n : ℤ) (a : α) : ‖a^n‖ ≤ ‖n‖ * ‖a‖ := +by rcases n.eq_coe_or_neg with ⟨n, rfl | rfl⟩; simpa using norm_pow_le_mul_norm n a + +@[to_additive nnnorm_zsmul_le] +lemma nnnorm_zpow_le_mul_norm (n : ℤ) (a : α) : ‖a^n‖₊ ≤ ‖n‖₊ * ‖a‖₊ := +by simpa only [← nnreal.coe_le_coe, nnreal.coe_mul] using norm_zpow_le_mul_norm n a + +end + namespace lipschitz_with variables [pseudo_emetric_space α] {K Kf Kg : ℝ≥0} {f g : α → E} diff --git a/src/analysis/normed/group/controlled_closure.lean b/src/analysis/normed/group/controlled_closure.lean new file mode 100644 index 0000000000000..535ac9b93a592 --- /dev/null +++ b/src/analysis/normed/group/controlled_closure.lean @@ -0,0 +1,121 @@ +/- +Copyright (c) 2021 Patrick Massot. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Patrick Massot +-/ +import analysis.normed.group.hom +import analysis.specific_limits.normed + +/-! # Extending a backward bound on a normed group homomorphism from a dense set + +Possible TODO (from the PR's review, https://github.com/leanprover-community/mathlib/pull/8498 ): +"This feels a lot like the second step in the proof of the Banach open mapping theorem +(`exists_preimage_norm_le`) ... wonder if it would be possible to refactor it using one of [the +lemmas in this file]." +-/ + +open filter finset +open_locale topological_space big_operators + +variables {G : Type*} [normed_add_comm_group G] [complete_space G] +variables {H : Type*} [normed_add_comm_group H] + +/-- Given `f : normed_add_group_hom G H` for some complete `G` and a subgroup `K` of `H`, if every +element `x` of `K` has a preimage under `f` whose norm is at most `C*‖x‖` then the same holds for +elements of the (topological) closure of `K` with constant `C+ε` instead of `C`, for any +positive `ε`. +-/ +lemma controlled_closure_of_complete {f : normed_add_group_hom G H} {K : add_subgroup H} + {C ε : ℝ} (hC : 0 < C) (hε : 0 < ε) (hyp : f.surjective_on_with K C) : + f.surjective_on_with K.topological_closure (C + ε) := +begin + rintros (h : H) (h_in : h ∈ K.topological_closure), + /- We first get rid of the easy case where `h = 0`.-/ + by_cases hyp_h : h = 0, + { rw hyp_h, + use 0, + simp }, + /- The desired preimage will be constructed as the sum of a series. Convergence of + the series will be guaranteed by completeness of `G`. We first write `h` as the sum + of a sequence `v` of elements of `K` which starts close to `h` and then quickly goes to zero. + The sequence `b` below quantifies this. -/ + set b : ℕ → ℝ := λ i, (1/2)^i*(ε*‖h‖/2)/C, + have b_pos : ∀ i, 0 < b i, + { intro i, + field_simp [b, hC], + exact div_pos (mul_pos hε (norm_pos_iff.mpr hyp_h)) + (mul_pos (by norm_num : (0 : ℝ) < 2^i*2) hC) }, + obtain ⟨v : ℕ → H, lim_v : tendsto (λ (n : ℕ), ∑ k in range (n + 1), v k) at_top (𝓝 h), + v_in : ∀ n, v n ∈ K, hv₀ : ‖v 0 - h‖ < b 0, hv : ∀ n > 0, ‖v n‖ < b n⟩ := + controlled_sum_of_mem_closure h_in b_pos, + /- The controlled surjectivity assumption on `f` allows to build preimages `u n` for all + elements `v n` of the `v` sequence.-/ + have : ∀ n, ∃ m' : G, f m' = v n ∧ ‖m'‖ ≤ C * ‖v n‖ := λ (n : ℕ), hyp (v n) (v_in n), + choose u hu hnorm_u using this, + /- The desired series `s` is then obtained by summing `u`. We then check our choice of + `b` ensures `s` is Cauchy. -/ + set s : ℕ → G := λ n, ∑ k in range (n+1), u k, + have : cauchy_seq s, + { apply normed_add_comm_group.cauchy_series_of_le_geometric'' (by norm_num) one_half_lt_one, + rintro n (hn : n ≥ 1), + calc ‖u n‖ ≤ C*‖v n‖ : hnorm_u n + ... ≤ C * b n : mul_le_mul_of_nonneg_left (hv _ $ nat.succ_le_iff.mp hn).le hC.le + ... = (1/2)^n * (ε * ‖h‖/2) : by simp [b, mul_div_cancel' _ hC.ne.symm] + ... = (ε * ‖h‖/2) * (1/2)^n : mul_comm _ _ }, + /- We now show that the limit `g` of `s` is the desired preimage. -/ + obtain ⟨g : G, hg⟩ := cauchy_seq_tendsto_of_complete this, + refine ⟨g, _, _⟩, + { /- We indeed get a preimage. First note: -/ + have : f ∘ s = λ n, ∑ k in range (n + 1), v k, + { ext n, + simp [map_sum, hu] }, + /- In the above equality, the left-hand-side converges to `f g` by continuity of `f` and + definition of `g` while the right-hand-side converges to `h` by construction of `v` so + `g` is indeed a preimage of `h`. -/ + rw ← this at lim_v, + exact tendsto_nhds_unique ((f.continuous.tendsto g).comp hg) lim_v }, + { /- Then we need to estimate the norm of `g`, using our careful choice of `b`. -/ + suffices : ∀ n, ‖s n‖ ≤ (C + ε) * ‖h‖, + from le_of_tendsto' (continuous_norm.continuous_at.tendsto.comp hg) this, + intros n, + have hnorm₀ : ‖u 0‖ ≤ C*b 0 + C*‖h‖, + { have := calc + ‖v 0‖ ≤ ‖h‖ + ‖v 0 - h‖ : norm_le_insert' _ _ + ... ≤ ‖h‖ + b 0 : by apply add_le_add_left hv₀.le, + calc ‖u 0‖ ≤ C*‖v 0‖ : hnorm_u 0 + ... ≤ C*(‖h‖ + b 0) : mul_le_mul_of_nonneg_left this hC.le + ... = C * b 0 + C * ‖h‖ : by rw [add_comm, mul_add] }, + have : ∑ k in range (n + 1), C * b k ≤ ε * ‖h‖ := calc + ∑ k in range (n + 1), C * b k = (∑ k in range (n + 1), (1 / 2) ^ k) * (ε * ‖h‖ / 2) : + by simp only [b, mul_div_cancel' _ hC.ne.symm, ← sum_mul] + ... ≤ 2 * (ε * ‖h‖ / 2) : mul_le_mul_of_nonneg_right (sum_geometric_two_le _) + (by nlinarith [hε, norm_nonneg h]) + ... = ε * ‖h‖ : mul_div_cancel' _ two_ne_zero, + calc ‖s n‖ ≤ ∑ k in range (n+1), ‖u k‖ : norm_sum_le _ _ + ... = ∑ k in range n, ‖u (k + 1)‖ + ‖u 0‖ : sum_range_succ' _ _ + ... ≤ ∑ k in range n, C*‖v (k + 1)‖ + ‖u 0‖ : add_le_add_right (sum_le_sum (λ _ _, hnorm_u _)) _ + ... ≤ ∑ k in range n, C*b (k+1) + (C*b 0 + C*‖h‖) : + add_le_add (sum_le_sum (λ k _, mul_le_mul_of_nonneg_left (hv _ k.succ_pos).le hC.le)) hnorm₀ + ... = ∑ k in range (n+1), C*b k + C*‖h‖ : by rw [← add_assoc, sum_range_succ'] + ... ≤ (C+ε)*‖h‖ : by { rw [add_comm, add_mul], apply add_le_add_left this } } +end + +/-- Given `f : normed_add_group_hom G H` for some complete `G`, if every element `x` of the image of +an isometric immersion `j : normed_add_group_hom K H` has a preimage under `f` whose norm is at most +`C*‖x‖` then the same holds for elements of the (topological) closure of this image with constant +`C+ε` instead of `C`, for any positive `ε`. +This is useful in particular if `j` is the inclusion of a normed group into its completion +(in this case the closure is the full target group). +-/ +lemma controlled_closure_range_of_complete {f : normed_add_group_hom G H} + {K : Type*} [seminormed_add_comm_group K] {j : normed_add_group_hom K H} (hj : ∀ x, ‖j x‖ = ‖x‖) + {C ε : ℝ} (hC : 0 < C) (hε : 0 < ε) (hyp : ∀ k, ∃ g, f g = j k ∧ ‖g‖ ≤ C*‖k‖) : + f.surjective_on_with j.range.topological_closure (C + ε) := +begin + replace hyp : ∀ h ∈ j.range, ∃ g, f g = h ∧ ‖g‖ ≤ C*‖h‖, + { intros h h_in, + rcases (j.mem_range _).mp h_in with ⟨k, rfl⟩, + rw hj, + exact hyp k }, + exact controlled_closure_of_complete hC hε hyp +end diff --git a/src/analysis/normed/group/hom.lean b/src/analysis/normed/group/hom.lean index 62bd8e1f499b3..4954266502108 100644 --- a/src/analysis/normed/group/hom.lean +++ b/src/analysis/normed/group/hom.lean @@ -3,8 +3,7 @@ Copyright (c) 2021 Johan Commelin. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin -/ - -import analysis.specific_limits.normed +import analysis.normed.group.basic /-! # Normed groups homomorphisms @@ -755,110 +754,3 @@ norm_lift_le _ _ _ hφ end equalizer end normed_add_group_hom - -section controlled_closure -open filter finset -open_locale topological_space -variables {G : Type*} [normed_add_comm_group G] [complete_space G] -variables {H : Type*} [normed_add_comm_group H] - -/-- Given `f : normed_add_group_hom G H` for some complete `G` and a subgroup `K` of `H`, if every -element `x` of `K` has a preimage under `f` whose norm is at most `C*‖x‖` then the same holds for -elements of the (topological) closure of `K` with constant `C+ε` instead of `C`, for any -positive `ε`. --/ -lemma controlled_closure_of_complete {f : normed_add_group_hom G H} {K : add_subgroup H} - {C ε : ℝ} (hC : 0 < C) (hε : 0 < ε) (hyp : f.surjective_on_with K C) : - f.surjective_on_with K.topological_closure (C + ε) := -begin - rintros (h : H) (h_in : h ∈ K.topological_closure), - /- We first get rid of the easy case where `h = 0`.-/ - by_cases hyp_h : h = 0, - { rw hyp_h, - use 0, - simp }, - /- The desired preimage will be constructed as the sum of a series. Convergence of - the series will be guaranteed by completeness of `G`. We first write `h` as the sum - of a sequence `v` of elements of `K` which starts close to `h` and then quickly goes to zero. - The sequence `b` below quantifies this. -/ - set b : ℕ → ℝ := λ i, (1/2)^i*(ε*‖h‖/2)/C, - have b_pos : ∀ i, 0 < b i, - { intro i, - field_simp [b, hC], - exact div_pos (mul_pos hε (norm_pos_iff.mpr hyp_h)) - (mul_pos (by norm_num : (0 : ℝ) < 2^i*2) hC) }, - obtain ⟨v : ℕ → H, lim_v : tendsto (λ (n : ℕ), ∑ k in range (n + 1), v k) at_top (𝓝 h), - v_in : ∀ n, v n ∈ K, hv₀ : ‖v 0 - h‖ < b 0, hv : ∀ n > 0, ‖v n‖ < b n⟩ := - controlled_sum_of_mem_closure h_in b_pos, - /- The controlled surjectivity assumption on `f` allows to build preimages `u n` for all - elements `v n` of the `v` sequence.-/ - have : ∀ n, ∃ m' : G, f m' = v n ∧ ‖m'‖ ≤ C * ‖v n‖ := λ (n : ℕ), hyp (v n) (v_in n), - choose u hu hnorm_u using this, - /- The desired series `s` is then obtained by summing `u`. We then check our choice of - `b` ensures `s` is Cauchy. -/ - set s : ℕ → G := λ n, ∑ k in range (n+1), u k, - have : cauchy_seq s, - { apply normed_add_comm_group.cauchy_series_of_le_geometric'' (by norm_num) one_half_lt_one, - rintro n (hn : n ≥ 1), - calc ‖u n‖ ≤ C*‖v n‖ : hnorm_u n - ... ≤ C * b n : mul_le_mul_of_nonneg_left (hv _ $ nat.succ_le_iff.mp hn).le hC.le - ... = (1/2)^n * (ε * ‖h‖/2) : by simp [b, mul_div_cancel' _ hC.ne.symm] - ... = (ε * ‖h‖/2) * (1/2)^n : mul_comm _ _ }, - /- We now show that the limit `g` of `s` is the desired preimage. -/ - obtain ⟨g : G, hg⟩ := cauchy_seq_tendsto_of_complete this, - refine ⟨g, _, _⟩, - { /- We indeed get a preimage. First note: -/ - have : f ∘ s = λ n, ∑ k in range (n + 1), v k, - { ext n, - simp [map_sum, hu] }, - /- In the above equality, the left-hand-side converges to `f g` by continuity of `f` and - definition of `g` while the right-hand-side converges to `h` by construction of `v` so - `g` is indeed a preimage of `h`. -/ - rw ← this at lim_v, - exact tendsto_nhds_unique ((f.continuous.tendsto g).comp hg) lim_v }, - { /- Then we need to estimate the norm of `g`, using our careful choice of `b`. -/ - suffices : ∀ n, ‖s n‖ ≤ (C + ε) * ‖h‖, - from le_of_tendsto' (continuous_norm.continuous_at.tendsto.comp hg) this, - intros n, - have hnorm₀ : ‖u 0‖ ≤ C*b 0 + C*‖h‖, - { have := calc - ‖v 0‖ ≤ ‖h‖ + ‖v 0 - h‖ : norm_le_insert' _ _ - ... ≤ ‖h‖ + b 0 : by apply add_le_add_left hv₀.le, - calc ‖u 0‖ ≤ C*‖v 0‖ : hnorm_u 0 - ... ≤ C*(‖h‖ + b 0) : mul_le_mul_of_nonneg_left this hC.le - ... = C * b 0 + C * ‖h‖ : by rw [add_comm, mul_add] }, - have : ∑ k in range (n + 1), C * b k ≤ ε * ‖h‖ := calc - ∑ k in range (n + 1), C * b k = (∑ k in range (n + 1), (1 / 2) ^ k) * (ε * ‖h‖ / 2) : - by simp only [b, mul_div_cancel' _ hC.ne.symm, ← sum_mul] - ... ≤ 2 * (ε * ‖h‖ / 2) : mul_le_mul_of_nonneg_right (sum_geometric_two_le _) - (by nlinarith [hε, norm_nonneg h]) - ... = ε * ‖h‖ : mul_div_cancel' _ two_ne_zero, - calc ‖s n‖ ≤ ∑ k in range (n+1), ‖u k‖ : norm_sum_le _ _ - ... = ∑ k in range n, ‖u (k + 1)‖ + ‖u 0‖ : sum_range_succ' _ _ - ... ≤ ∑ k in range n, C*‖v (k + 1)‖ + ‖u 0‖ : add_le_add_right (sum_le_sum (λ _ _, hnorm_u _)) _ - ... ≤ ∑ k in range n, C*b (k+1) + (C*b 0 + C*‖h‖) : - add_le_add (sum_le_sum (λ k _, mul_le_mul_of_nonneg_left (hv _ k.succ_pos).le hC.le)) hnorm₀ - ... = ∑ k in range (n+1), C*b k + C*‖h‖ : by rw [← add_assoc, sum_range_succ'] - ... ≤ (C+ε)*‖h‖ : by { rw [add_comm, add_mul], apply add_le_add_left this } } -end - -/-- Given `f : normed_add_group_hom G H` for some complete `G`, if every element `x` of the image of -an isometric immersion `j : normed_add_group_hom K H` has a preimage under `f` whose norm is at most -`C*‖x‖` then the same holds for elements of the (topological) closure of this image with constant -`C+ε` instead of `C`, for any positive `ε`. -This is useful in particular if `j` is the inclusion of a normed group into its completion -(in this case the closure is the full target group). --/ -lemma controlled_closure_range_of_complete {f : normed_add_group_hom G H} - {K : Type*} [seminormed_add_comm_group K] {j : normed_add_group_hom K H} (hj : ∀ x, ‖j x‖ = ‖x‖) - {C ε : ℝ} (hC : 0 < C) (hε : 0 < ε) (hyp : ∀ k, ∃ g, f g = j k ∧ ‖g‖ ≤ C*‖k‖) : - f.surjective_on_with j.range.topological_closure (C + ε) := -begin - replace hyp : ∀ h ∈ j.range, ∃ g, f g = h ∧ ‖g‖ ≤ C*‖h‖, - { intros h h_in, - rcases (j.mem_range _).mp h_in with ⟨k, rfl⟩, - rw hj, - exact hyp k }, - exact controlled_closure_of_complete hC hε hyp -end -end controlled_closure diff --git a/src/analysis/normed/group/quotient.lean b/src/analysis/normed/group/quotient.lean index 836809a7fec30..963dda9704818 100644 --- a/src/analysis/normed/group/quotient.lean +++ b/src/analysis/normed/group/quotient.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Patrick Massot. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Patrick Massot, Riccardo Brasca -/ +import analysis.normed_space.basic import analysis.normed.group.hom /-! diff --git a/src/analysis/special_functions/exp.lean b/src/analysis/special_functions/exp.lean index 7960020f35ea6..553a2a3e86118 100644 --- a/src/analysis/special_functions/exp.lean +++ b/src/analysis/special_functions/exp.lean @@ -3,9 +3,10 @@ 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.asymptotics.theta import analysis.complex.basic +import analysis.specific_limits.normed import data.complex.exponential -import analysis.asymptotics.theta /-! # Complex and real exponential From b84da9c84486f209081466c8045d4a646828672e Mon Sep 17 00:00:00 2001 From: Devon Tuma Date: Thu, 19 Jan 2023 10:30:47 +0000 Subject: [PATCH 0295/1291] feat(topology/algebra/infinite_sum): Generalized lemmas for `add_comm_monoid` rather than `add_comm_group` (#17035) This PR adds generalized (but weaker) versions of `has_sum.update`, `has_sum.ite_eq_extract`, and `tsum_ite_eq_extract` when the co-domain is a monoid rather than a group. The main use of this is with `nnreal` and `ennreal`, to write things like a `tsum` over an `option` type as a sum of the `none` and `some` cases separately. --- src/topology/algebra/infinite_sum.lean | 57 ++++++++++++++++++++++++-- src/topology/instances/ennreal.lean | 11 +++++ 2 files changed, 64 insertions(+), 4 deletions(-) diff --git a/src/topology/algebra/infinite_sum.lean b/src/topology/algebra/infinite_sum.lean index 2ba8e2083616a..7eaa1e5dc1354 100644 --- a/src/topology/algebra/infinite_sum.lean +++ b/src/topology/algebra/infinite_sum.lean @@ -405,6 +405,35 @@ lemma has_sum.sigma_of_has_sum [t3_space α] {γ : β → Type*} {f : (Σ b:β, has_sum f a := by simpa [(hf'.has_sum.sigma hf).unique ha] using hf'.has_sum +/-- Version of `has_sum.update` for `add_comm_monoid` rather than `add_comm_group`. +Rather than showing that `f.update` has a specific sum in terms of `has_sum`, +it gives a relationship between the sums of `f` and `f.update` given that both exist. -/ +lemma has_sum.update' {α β : Type*} [topological_space α] [add_comm_monoid α] [t2_space α] + [has_continuous_add α] {f : β → α} {a a' : α} (hf : has_sum f a) + (b : β) (x : α) (hf' : has_sum (f.update b x) a') : a + x = a' + f b := +begin + have : ∀ b', f b' + ite (b' = b) x 0 = f.update b x b' + ite (b' = b) (f b) 0, + { intro b', + split_ifs with hb', + { simpa only [function.update_apply, hb', eq_self_iff_true] using add_comm (f b) x }, + { simp only [function.update_apply, hb', if_false] } }, + have h := hf.add ((has_sum_ite_eq b x)), + simp_rw this at h, + exact has_sum.unique h (hf'.add (has_sum_ite_eq b (f b))) +end + +/-- Version of `has_sum_ite_sub_has_sum` for `add_comm_monoid` rather than `add_comm_group`. +Rather than showing that the `ite` expression has a specific sum in terms of `has_sum`, +it gives a relationship between the sums of `f` and `ite (n = b) 0 (f n)` given that both exist. -/ +lemma eq_add_of_has_sum_ite {α β : Type*} [topological_space α] [add_comm_monoid α] + [t2_space α] [has_continuous_add α] {f : β → α} {a : α} (hf : has_sum f a) (b : β) (a' : α) + (hf' : has_sum (λ n, ite (n = b) 0 (f n)) a') : a = a' + f b := +begin + refine (add_zero a).symm.trans (hf.update' b 0 _), + convert hf', + exact funext (f.update_apply b 0), +end + end has_sum section tsum @@ -446,6 +475,13 @@ lemma tsum_eq_sum {f : β → α} {s : finset β} (hf : ∀b∉s, f b = 0) : ∑' b, f b = ∑ b in s, f b := (has_sum_sum_of_ne_finset_zero hf).tsum_eq +lemma sum_eq_tsum_indicator (f : β → α) (s : finset β) : + ∑ x in s, f x = ∑' x, set.indicator ↑s f x := +have ∀ x ∉ s, set.indicator ↑s f x = 0, +from λ x hx, set.indicator_apply_eq_zero.2 (λ hx', (hx $ finset.mem_coe.1 hx').elim), +(finset.sum_congr rfl (λ x hx, (set.indicator_apply_eq_self.2 $ + λ hx', (hx' $ finset.mem_coe.2 hx).elim).symm)).trans (tsum_eq_sum this).symm + lemma tsum_congr {α β : Type*} [add_comm_monoid α] [topological_space α] {f g : β → α} (hfg : ∀ b, f b = g b) : ∑' b, f b = ∑' b, g b := congr_arg tsum (funext hfg) @@ -569,6 +605,19 @@ lemma tsum_sum {f : γ → β → α} {s : finset γ} (hf : ∀i∈s, summable ( ∑'b, ∑ i in s, f i b = ∑ i in s, ∑'b, f i b := (has_sum_sum $ assume i hi, (hf i hi).has_sum).tsum_eq +/-- Version of `tsum_eq_add_tsum_ite` for `add_comm_monoid` rather than `add_comm_group`. +Requires a different convergence assumption involving `function.update`. -/ +lemma tsum_eq_add_tsum_ite' {f : β → α} (b : β) (hf : summable (f.update b 0)) : + ∑' x, f x = f b + ∑' x, ite (x = b) 0 (f x) := +calc ∑' x, f x = ∑' x, ((ite (x = b) (f x) 0) + (f.update b 0 x)) : + tsum_congr (λ n, by split_ifs; simp [function.update_apply, h]) + ... = ∑' x, ite (x = b) (f x) 0 + ∑' x, f.update b 0 x : + tsum_add ⟨ite (b = b) (f b) 0, has_sum_single b (λ b hb, if_neg hb)⟩ (hf) + ... = (ite (b = b) (f b) 0) + ∑' x, f.update b 0 x : + by { congr, exact (tsum_eq_single b (λ b' hb', if_neg hb')) } + ... = f b + ∑' x, ite (x = b) 0 (f x) : + by simp only [function.update, eq_self_iff_true, if_true, eq_rec_constant, dite_eq_ite] + variables [add_comm_monoid δ] [topological_space δ] [t3_space δ] [has_continuous_add δ] lemma tsum_sigma' {γ : β → Type*} {f : (Σb:β, γ b) → δ} (h₁ : ∀b, summable (λc, f ⟨b, c⟩)) @@ -798,7 +847,7 @@ lemma set.finite.summable_compl_iff {s : set β} (hs : s.finite) : summable (f ∘ coe : sᶜ → α) ↔ summable f := (hs.summable f).summable_compl_iff -lemma has_sum_ite_eq_extract [decidable_eq β] (hf : has_sum f a) (b : β) : +lemma has_sum_ite_sub_has_sum [decidable_eq β] (hf : has_sum f a) (b : β) : has_sum (λ n, ite (n = b) 0 (f n)) (a - f b) := begin convert hf.update b 0 using 1, @@ -824,12 +873,12 @@ lemma sum_add_tsum_compl {s : finset β} (hf : summable f) : ((s.has_sum f).add_compl (s.summable_compl_iff.2 hf).has_sum).tsum_eq.symm /-- Let `f : β → α` be a sequence with summable series and let `b ∈ β` be an index. -Lemma `tsum_ite_eq_extract` writes `Σ f n` as the sum of `f b` plus the series of the +Lemma `tsum_eq_add_tsum_ite` writes `Σ f n` as the sum of `f b` plus the series of the remaining terms. -/ -lemma tsum_ite_eq_extract [decidable_eq β] (hf : summable f) (b : β) : +lemma tsum_eq_add_tsum_ite [decidable_eq β] (hf : summable f) (b : β) : ∑' n, f n = f b + ∑' n, ite (n = b) 0 (f n) := begin - rw (has_sum_ite_eq_extract hf.has_sum b).tsum_eq, + rw (has_sum_ite_sub_has_sum hf.has_sum b).tsum_eq, exact (add_sub_cancel'_right _ _).symm, end diff --git a/src/topology/instances/ennreal.lean b/src/topology/instances/ennreal.lean index 8d5af6c912c16..5bf8b42b77410 100644 --- a/src/topology/instances/ennreal.lean +++ b/src/topology/instances/ennreal.lean @@ -964,6 +964,9 @@ begin exact tsum_bUnion_le _ _ _ end +lemma tsum_eq_add_tsum_ite {f : β → ℝ≥0∞} (b : β) : ∑' x, f x = f b + ∑' x, ite (x = b) 0 (f x) := +tsum_eq_add_tsum_ite' b ennreal.summable + lemma tsum_add_one_eq_top {f : ℕ → ℝ≥0∞} (hf : ∑' n, f n = ∞) (hf0 : f 0 ≠ ∞) : ∑' n, f (n + 1) = ∞ := begin @@ -1205,6 +1208,14 @@ lemma tsum_pos {g : α → ℝ≥0} (hg : summable g) (i : α) (hi : 0 < g i) : 0 < ∑' b, g b := by { rw ← tsum_zero, exact tsum_lt_tsum (λ a, zero_le _) hi hg } +lemma tsum_eq_add_tsum_ite {f : α → ℝ≥0} (hf : summable f) (i : α) : + ∑' x, f x = f i + ∑' x, ite (x = i) 0 (f x) := +begin + refine tsum_eq_add_tsum_ite' i (nnreal.summable_of_le (λ i', _) hf), + rw [function.update_apply], + split_ifs; simp only [zero_le', le_rfl] +end + end nnreal namespace ennreal From e94cd75f7f28fc7978d71eb4e6e41925ae9a858b Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 19 Jan 2023 13:45:47 +0000 Subject: [PATCH 0296/1291] chore(analysis/complex/cauchy_integral): squeeze simp (#18225) This speeds up the proof by a factor of 3 (from 17s to 5s on gitpod) --- src/analysis/complex/cauchy_integral.lean | 5 ++++- 1 file changed, 4 insertions(+), 1 deletion(-) diff --git a/src/analysis/complex/cauchy_integral.lean b/src/analysis/complex/cauchy_integral.lean index 2cbc77d447e99..b4c6d6e96f5ee 100644 --- a/src/analysis/complex/cauchy_integral.lean +++ b/src/analysis/complex/cauchy_integral.lean @@ -176,7 +176,10 @@ begin set F : (ℝ × ℝ) → E := f ∘ e, set F' : (ℝ × ℝ) → (ℝ × ℝ) →L[ℝ] E := λ p, (f' (e p)).comp (e : (ℝ × ℝ) →L[ℝ] ℂ), have hF' : ∀ p : ℝ × ℝ, (-(I • F' p)) (1, 0) + F' p (0, 1) = -(I • f' (e p) 1 - f' (e p) I), - { rintro ⟨x, y⟩, simp [F', he₁, he₂, ← sub_eq_neg_add], }, + { rintro ⟨x, y⟩, + simp only [continuous_linear_map.neg_apply, continuous_linear_map.smul_apply, F', + continuous_linear_map.comp_apply, continuous_linear_equiv.coe_coe, he₁, he₂, + neg_add_eq_sub, neg_sub], }, set R : set (ℝ × ℝ) := [z.re, w.re] ×ˢ [w.im, z.im], set t : set (ℝ × ℝ) := e ⁻¹' s, rw [uIcc_comm z.im] at Hc Hi, rw [min_comm z.im, max_comm z.im] at Hd, From 9292ecc21fc33477bfb6b5a91d614b52b1889f38 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 19 Jan 2023 15:23:46 +0000 Subject: [PATCH 0297/1291] feat(analysis/special_functions/pow): interaction between `cpow` and `inv`/`conj` (#18137) --- src/analysis/special_functions/pow.lean | 37 +++++++++++++++++++++++++ 1 file changed, 37 insertions(+) diff --git a/src/analysis/special_functions/pow.lean b/src/analysis/special_functions/pow.lean index 4fd66dc7d3314..ddcfb486c74dd 100644 --- a/src/analysis/special_functions/pow.lean +++ b/src/analysis/special_functions/pow.lean @@ -21,6 +21,7 @@ We also prove basic properties of these functions. noncomputable theory open_locale classical real topological_space nnreal ennreal filter big_operators asymptotics +open_locale complex_conjugate open filter finset set namespace complex @@ -493,6 +494,42 @@ begin { exact abs_cpow_eq_rpow_re_of_pos hlt y } end +lemma inv_cpow_eq_ite (x : ℂ) (n : ℂ) : + x⁻¹ ^ n = if x.arg = π then conj (x ^ conj n)⁻¹ else (x ^ n)⁻¹ := +begin + simp_rw [complex.cpow_def, log_inv_eq_ite, inv_eq_zero, map_eq_zero, ite_mul, neg_mul, + is_R_or_C.conj_inv, apply_ite conj, apply_ite exp, apply_ite has_inv.inv, map_zero, map_one, + exp_neg, inv_one, inv_zero, ←exp_conj, map_mul, conj_conj], + split_ifs with hx hn ha ha; refl, +end + +lemma inv_cpow (x : ℂ) (n : ℂ) (hx : x.arg ≠ π) : x⁻¹ ^ n = (x ^ n)⁻¹ := +by rw [inv_cpow_eq_ite, if_neg hx] + +/-- `complex.inv_cpow_eq_ite` with the `ite` on the other side. -/ +lemma inv_cpow_eq_ite' (x : ℂ) (n : ℂ) : + (x ^ n)⁻¹ = if x.arg = π then conj (x⁻¹ ^ conj n) else x⁻¹ ^ n := +begin + rw [inv_cpow_eq_ite, apply_ite conj, conj_conj, conj_conj], + split_ifs, + { refl }, + { rw inv_cpow _ _ h } +end + +lemma conj_cpow_eq_ite (x : ℂ) (n : ℂ) : + conj x ^ n = if x.arg = π then x ^ n else conj (x ^ conj n) := +begin + simp_rw [cpow_def, map_eq_zero, apply_ite conj, map_one, map_zero, ←exp_conj, map_mul, + conj_conj, log_conj_eq_ite], + split_ifs with hcx hn hx; refl +end + +lemma conj_cpow (x : ℂ) (n : ℂ) (hx : x.arg ≠ π) : conj x ^ n = conj (x ^ conj n) := +by rw [conj_cpow_eq_ite, if_neg hx] + +lemma cpow_conj (x : ℂ) (n : ℂ) (hx : x.arg ≠ π) : x ^ conj n = conj (conj x ^ n) := +by rw [conj_cpow _ _ hx, conj_conj] + end complex namespace real From 2609ad04e2b77aceb2200e273daf11cc65856fda Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Thu, 19 Jan 2023 15:23:48 +0000 Subject: [PATCH 0298/1291] chore(data/nat/digits): golf, use seemingly weaker assumptions (#18203) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit * Golf. * Assume `n ≠ 0` instead of `0 < n` or `1 ≤ n`. * Assume `1 < n` instead of `2 ≤ n`. * Add `nat.exists_eq_add_of_le'` (forward-ported in leanprover-community/mathlib4#1662). --- src/data/nat/basic.lean | 20 +++--- src/data/nat/digits.lean | 139 ++++++++++++++++----------------------- 2 files changed, 63 insertions(+), 96 deletions(-) diff --git a/src/data/nat/basic.lean b/src/data/nat/basic.lean index 4b5a0b8a04040..d8f47d7509cdd 100644 --- a/src/data/nat/basic.lean +++ b/src/data/nat/basic.lean @@ -210,18 +210,14 @@ by { rw ←not_iff_not, push_neg, exact forall_lt_succ } @[simp] theorem add_def {a b : ℕ} : nat.add a b = a + b := rfl @[simp] theorem mul_def {a b : ℕ} : nat.mul a b = a * b := rfl -lemma exists_eq_add_of_le : ∀ {m n : ℕ}, m ≤ n → ∃ k : ℕ, n = m + k -| 0 0 h := ⟨0, by simp⟩ -| 0 (n+1) h := ⟨n+1, by simp⟩ -| (m+1) (n+1) h := - let ⟨k, hk⟩ := exists_eq_add_of_le (nat.le_of_succ_le_succ h) in - ⟨k, by simp [hk, add_comm, add_left_comm]⟩ - -lemma exists_eq_add_of_lt : ∀ {m n : ℕ}, m < n → ∃ k : ℕ, n = m + k + 1 -| 0 0 h := false.elim $ lt_irrefl _ h -| 0 (n+1) h := ⟨n, by simp⟩ -| (m+1) (n+1) h := let ⟨k, hk⟩ := exists_eq_add_of_le (nat.le_of_succ_le_succ h) in - ⟨k, by simp [hk]⟩ +lemma exists_eq_add_of_le (h : m ≤ n) : ∃ k : ℕ, n = m + k := +⟨n - m, (nat.add_sub_of_le h).symm⟩ + +lemma exists_eq_add_of_le' (h : m ≤ n) : ∃ k : ℕ, n = k + m := +⟨n - m, (nat.sub_add_cancel h).symm⟩ + +lemma exists_eq_add_of_lt (h : m < n) : ∃ k : ℕ, n = m + k + 1 := +⟨n - (m + 1), by rw [add_right_comm, nat.add_sub_of_le h]⟩ /-! ### `pred` -/ diff --git a/src/data/nat/digits.lean b/src/data/nat/digits.lean index ba1681b9897fb..248bf2a028399 100644 --- a/src/data/nat/digits.lean +++ b/src/data/nat/digits.lean @@ -79,8 +79,8 @@ by rcases b with _|⟨_|⟨_⟩⟩; simp [digits, digits_aux_0, digits_aux_1] @[simp] lemma digits_zero_succ (n : ℕ) : digits 0 (n.succ) = [n+1] := rfl -theorem digits_zero_succ' : ∀ {n : ℕ} (w : 0 < n), digits 0 n = [n] -| 0 h := absurd h dec_trivial +theorem digits_zero_succ' : ∀ {n : ℕ}, n ≠ 0 → digits 0 n = [n] +| 0 h := (h rfl).elim | (n+1) _ := rfl @[simp] lemma digits_one (n : ℕ) : digits 1 n = list.replicate n 1 := rfl @@ -91,40 +91,31 @@ theorem digits_zero_succ' : ∀ {n : ℕ} (w : 0 < n), digits 0 n = [n] digits (b+2) (n+1) = (((n+1) % (b+2)) :: digits (b+2) ((n+1) / (b+2))) := by { rw [digits, digits_aux_def], exact succ_pos n } -theorem digits_def' : ∀ {b : ℕ} (h : 2 ≤ b) {n : ℕ} (w : 0 < n), +theorem digits_def' : ∀ {b : ℕ} (h : 1 < b) {n : ℕ} (w : 0 < n), digits b n = n % b :: digits b (n/b) | 0 h := absurd h dec_trivial | 1 h := absurd h dec_trivial | (b+2) h := digits_aux_def _ _ -@[simp] -lemma digits_of_lt (b x : ℕ) (w₁ : 0 < x) (w₂ : x < b) : digits b x = [x] := +@[simp] lemma digits_of_lt (b x : ℕ) (hx : x ≠ 0) (hxb : x < b) : digits b x = [x] := begin - cases b, - { cases w₂ }, - { cases b, - { interval_cases x, }, - { cases x, - { cases w₁, }, - { rw [digits_add_two_add_one, nat.div_eq_of_lt w₂, digits_zero, nat.mod_eq_of_lt w₂] } } } + rcases exists_eq_succ_of_ne_zero hx with ⟨x, rfl⟩, + rcases exists_eq_add_of_le' ((nat.le_add_left 1 x).trans_lt hxb) with ⟨b, rfl⟩, + rw [digits_add_two_add_one, div_eq_of_lt hxb, digits_zero, mod_eq_of_lt hxb] end -lemma digits_add (b : ℕ) (h : 2 ≤ b) (x y : ℕ) (w : x < b) (w' : 0 < x ∨ 0 < y) : +lemma digits_add (b : ℕ) (h : 1 < b) (x y : ℕ) (hxb : x < b) (hxy : x ≠ 0 ∨ y ≠ 0) : digits b (x + b * y) = x :: digits b y := begin - cases b, - { cases h, }, - { cases b, - { norm_num at h, }, - { cases y, - { norm_num at w', - simp [w, w'], }, - dsimp [digits], - rw digits_aux_def, - { congr, - { simp [nat.add_mod, nat.mod_eq_of_lt w], }, - { simp [mul_comm (b+2), nat.add_mul_div_right, nat.div_eq_of_lt w], } }, - { apply nat.succ_pos, }, }, }, + rcases exists_eq_add_of_le' h with ⟨b, rfl : _ = _ + 2⟩, + cases y, + { simp [hxb, hxy.resolve_right (absurd rfl)] }, + dsimp [digits], + rw digits_aux_def, + { congr, + { simp [nat.add_mod, mod_eq_of_lt hxb], }, + { simp [add_mul_div_left, div_eq_of_lt hxb] } }, + { apply nat.succ_pos } end /-- @@ -201,22 +192,14 @@ begin { dsimp [of_digits], push_cast } end -lemma digits_zero_of_eq_zero {b : ℕ} (h : 1 ≤ b) {L : list ℕ} (w : of_digits b L = 0) : - ∀ l ∈ L, l = 0 := -begin - induction L with d L ih, - { intros l m, - cases m, }, - { intros l m, - dsimp [of_digits] at w, - rcases m with ⟨rfl⟩, - { apply nat.eq_zero_of_add_eq_zero_right w }, - { exact ih (mul_right_injective₀ (pos_iff_ne_zero.1 h) - (nat.eq_zero_of_add_eq_zero_left w)) _ m, }, } -end +lemma digits_zero_of_eq_zero {b : ℕ} (h : b ≠ 0) : + ∀ {L : list ℕ} (h0 : of_digits b L = 0) (l ∈ L), l = 0 +| (a :: L) h0 l (or.inl rfl) := nat.eq_zero_of_add_eq_zero_right h0 +| (a :: L) h0 l (or.inr hL) := + digits_zero_of_eq_zero (mul_right_injective₀ h (nat.eq_zero_of_add_eq_zero_left h0)) _ hL lemma digits_of_digits - (b : ℕ) (h : 2 ≤ b) (L : list ℕ) + (b : ℕ) (h : 1 < b) (L : list ℕ) (w₁ : ∀ l ∈ L, l < b) (w₂ : ∀ (h : L ≠ []), L.last h ≠ 0) : digits b (of_digits b L) = L := begin @@ -233,17 +216,13 @@ begin { exact w₁ d (list.mem_cons_self _ _) }, { by_cases h' : L = [], { rcases h' with rfl, - simp at w₂, left, - apply nat.pos_of_ne_zero, - exact w₂ }, + simpa using w₂ }, { right, - apply nat.pos_of_ne_zero, contrapose! w₂, - apply digits_zero_of_eq_zero _ w₂, - { rw list.last_cons h', - exact list.last_mem h', }, - { exact le_of_lt h, }, }, }, }, + refine digits_zero_of_eq_zero h.ne_bot w₂ _ _, + rw list.last_cons h', + exact list.last_mem h' } } } end lemma of_digits_digits (b n : ℕ) : of_digits b (digits b n) = n := @@ -295,7 +274,7 @@ end lemma digits_ne_nil_iff_ne_zero {b n : ℕ} : digits b n ≠ [] ↔ n ≠ 0 := not_congr digits_eq_nil_iff_eq_zero -lemma digits_eq_cons_digits_div {b n : ℕ} (h : 2 ≤ b) (w : 0 < n) : +lemma digits_eq_cons_digits_div {b n : ℕ} (h : 1 < b) (w : n ≠ 0) : digits b n = ((n % b) :: digits b (n / b)) := begin rcases b with _|_|b, @@ -306,12 +285,12 @@ begin simp, end -lemma digits_last {b : ℕ} (m : ℕ) (h : 2 ≤ b) (p q) : +lemma digits_last {b : ℕ} (m : ℕ) (h : 1 < b) (p q) : (digits b m).last p = (digits b (m/b)).last q := begin by_cases hm : m = 0, { simp [hm], }, - simp only [digits_eq_cons_digits_div h (nat.pos_of_ne_zero hm)], + simp only [digits_eq_cons_digits_div h hm], rw list.last_cons, end @@ -322,21 +301,20 @@ function.left_inverse.injective (of_digits_digits b) b.digits n = b.digits m ↔ n = m := (digits.injective b).eq_iff -lemma digits_len (b n : ℕ) (hb : 2 ≤ b) (hn : 0 < n) : +lemma digits_len (b n : ℕ) (hb : 1 < b) (hn : n ≠ 0) : (b.digits n).length = b.log n + 1 := begin induction n using nat.strong_induction_on with n IH, rw [digits_eq_cons_digits_div hb hn, list.length], - cases (n / b).eq_zero_or_pos with h h, - { have posb : 0 < b := zero_lt_two.trans_le hb, - simp [h, log_eq_zero_iff, ←nat.div_eq_zero_iff posb] }, + by_cases h : n / b = 0, + { have hb0 : b ≠ 0 := (nat.succ_le_iff.1 hb).ne_bot, + simp [h, log_eq_zero_iff, ← nat.div_eq_zero_iff hb0.bot_lt] }, { have hb' : 1 < b := one_lt_two.trans_le hb, - have : n / b < n := div_lt_self hn hb', + have : n / b < n := div_lt_self (nat.pos_of_ne_zero hn) hb', rw [IH _ this h, log_div_base, tsub_add_cancel_of_le], - rw [succ_le_iff], - refine log_pos hb' _, + refine nat.succ_le_of_lt (log_pos hb' _), contrapose! h, - rw div_eq_of_lt h } + exact div_eq_of_lt h } end lemma last_digit_ne_zero (b : ℕ) {m : ℕ} (hm : m ≠ 0) : @@ -351,12 +329,10 @@ begin revert hm, apply nat.strong_induction_on m, intros n IH hn, - have hnpos : 0 < n := nat.pos_of_ne_zero hn, by_cases hnb : n < b + 2, - { simp_rw [digits_of_lt b.succ.succ n hnpos hnb], - exact pos_iff_ne_zero.mp hnpos }, + { simpa only [digits_of_lt (b + 2) n hn hnb] }, { rw digits_last n (show 2 ≤ b + 2, from dec_trivial), - refine IH _ (nat.div_lt_self hnpos dec_trivial) _, + refine IH _ (nat.div_lt_self hn.bot_lt dec_trivial) _, { rw ←pos_iff_ne_zero, exact nat.div_pos (le_of_not_lt hnb) dec_trivial } }, end @@ -375,7 +351,7 @@ begin end /-- The digits in the base b expansion of n are all less than b, if b ≥ 2 -/ -lemma digits_lt_base {b m d : ℕ} (hb : 2 ≤ b) (hd : d ∈ digits b m) : d < b := +lemma digits_lt_base {b m d : ℕ} (hb : 1 < b) (hd : d ∈ digits b m) : d < b := begin rcases b with _ | _ | b; try {linarith}, exact digits_lt_base' hd, @@ -397,8 +373,8 @@ begin exact hl hd (list.mem_cons_self _ _) } end -/-- an n-digit number in base b is less than b^n if b ≥ 2 -/ -lemma of_digits_lt_base_pow_length {b : ℕ} {l : list ℕ} (hb : 2 ≤ b) (hl : ∀ x ∈ l, x < b) : +/-- an n-digit number in base b is less than b^n if b > 1 -/ +lemma of_digits_lt_base_pow_length {b : ℕ} {l : list ℕ} (hb : 1 < b) (hl : ∀ x ∈ l, x < b) : of_digits b l < b^l.length := begin rcases b with _ | _ | b; try { linarith }, @@ -413,7 +389,7 @@ begin end /-- Any number m is less than b^(number of digits in the base b representation of m) -/ -lemma lt_base_pow_length_digits {b m : ℕ} (hb : 2 ≤ b) : m < b^(digits b m).length := +lemma lt_base_pow_length_digits {b m : ℕ} (hb : 1 < b) : m < b^(digits b m).length := begin rcases b with _ | _ | b; try { linarith }, exact lt_base_pow_length_digits', @@ -425,13 +401,10 @@ by rw [of_digits_append, of_digits_digits, of_digits_digits] lemma digits_len_le_digits_len_succ (b n : ℕ) : (digits b n).length ≤ (digits b (n + 1)).length := begin - rcases n.eq_zero_or_pos with rfl|hn, + rcases decidable.eq_or_ne n 0 with rfl|hn, { simp }, - cases lt_or_le b 2 with hb hb, - { rcases b with _|_|b, - { simp [digits_zero_succ', hn] }, - { simp, }, - { simpa [succ_lt_succ_iff] using hb } }, + cases le_or_lt b 1 with hb hb, + { interval_cases b; simp [digits_zero_succ', hn] }, simpa [digits_len, hb, hn] using log_mono_right (le_succ _) end @@ -466,7 +439,7 @@ end Any non-zero natural number `m` is greater than b^((number of digits in the base b representation of m) - 1) -/ -lemma base_pow_length_digits_le (b m : ℕ) (hb : 2 ≤ b): m ≠ 0 → b ^ ((digits b m).length) ≤ b * m := +lemma base_pow_length_digits_le (b m : ℕ) (hb : 1 < b): m ≠ 0 → b ^ ((digits b m).length) ≤ b * m := begin rcases b with _ | _ | b; try { linarith }, exact base_pow_length_digits_le' b m, @@ -480,11 +453,10 @@ begin { simp, }, rw bits_append_bit _ _ (λ hn, absurd hn h), cases b, - { rw digits_def' (le_refl 2), + { rw digits_def' one_lt_two, { simpa [nat.bit, nat.bit0_val n], }, { simpa [pos_iff_ne_zero, bit_eq_zero_iff], }, }, - { simpa [nat.bit, nat.bit1_val n, add_comm, digits_add 2 le_rfl 1 n (by norm_num) - (by norm_num)] }, + { simpa [nat.bit, nat.bit1_val n, add_comm, digits_add 2 one_lt_two 1 n] }, end @@ -629,8 +601,8 @@ theorem digits_succ (b n m r l) (e : r + b * m = n) (hr : r < b) - (h : nat.digits b m = l ∧ 2 ≤ b ∧ 0 < m) : - nat.digits b n = r :: l ∧ 2 ≤ b ∧ 0 < n := + (h : nat.digits b m = l ∧ 1 < b ∧ 0 < m) : + nat.digits b n = r :: l ∧ 1 < b ∧ 0 < n := begin rcases h with ⟨h, b2, m0⟩, have b0 : 0 < b := by linarith, @@ -642,12 +614,12 @@ end theorem digits_one (b n) (n0 : 0 < n) (nb : n < b) : - nat.digits b n = [n] ∧ 2 ≤ b ∧ 0 < n := + nat.digits b n = [n] ∧ 1 < b ∧ 0 < n := begin - have b2 : 2 ≤ b := by linarith, + have b2 : 1 < b := by linarith, refine ⟨_, b2, n0⟩, rw [nat.digits_def' b2 n0, nat.mod_eq_of_lt nb, - (nat.div_eq_zero_iff (by linarith : 0 < b)).2 nb, nat.digits_zero], + (nat.div_eq_zero_iff ((zero_le n).trans_lt nb)).2 nb, nat.digits_zero], end open tactic @@ -685,11 +657,10 @@ example : nat.digits 10 123 = [3,2,1] := by norm_num if n = 0 then return (`([] : list ℕ), `(nat.digits_zero %%eb)) else if b = 0 then do ic ← mk_instance_cache `(ℕ), - (_, pn0) ← norm_num.prove_pos ic en, + (_, pn0) ← norm_num.prove_ne_zero' ic en, return (`([%%en] : list ℕ), `(@nat.digits_zero_succ' %%en %%pn0)) else if b = 1 then do ic ← mk_instance_cache `(ℕ), - (_, pn0) ← norm_num.prove_pos ic en, s ← simp_lemmas.add_simp simp_lemmas.mk `list.replicate, (rhs, p2, _) ← simplify s [] `(list.replicate %%en 1), p ← mk_eq_trans `(nat.digits_one %%en) p2, From 78314d08d707a6338079f00094bbdb90bf11fc41 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 19 Jan 2023 15:23:50 +0000 Subject: [PATCH 0299/1291] chore(data/fintype/vector, logic/equiv/list): split (#18226) `array` and `d_array` do not exist in Lean 4; it's easier for porting if we put the stuff about those types in separate files. --- src/computability/primrec.lean | 1 + src/data/array/lemmas.lean | 29 ------------------- src/data/fintype/array.lean | 20 +++++++++++++ src/data/fintype/vector.lean | 10 +------ src/logic/equiv/array.lean | 52 ++++++++++++++++++++++++++++++++++ src/logic/equiv/list.lean | 10 +------ 6 files changed, 75 insertions(+), 47 deletions(-) create mode 100644 src/data/fintype/array.lean create mode 100644 src/logic/equiv/array.lean diff --git a/src/computability/primrec.lean b/src/computability/primrec.lean index 669d748677efa..e6b4673132d2f 100644 --- a/src/computability/primrec.lean +++ b/src/computability/primrec.lean @@ -3,6 +3,7 @@ Copyright (c) 2018 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ +import logic.equiv.array import logic.equiv.list import logic.function.iterate diff --git a/src/data/array/lemmas.lean b/src/data/array/lemmas.lean index 15745bf153213..2f0f22eff241c 100644 --- a/src/data/array/lemmas.lean +++ b/src/data/array/lemmas.lean @@ -270,32 +270,3 @@ read_foreach end map₂ end array - -namespace equiv - -/-- The natural equivalence between length-`n` heterogeneous arrays -and dependent functions from `fin n`. -/ -def d_array_equiv_fin {n : ℕ} (α : fin n → Type*) : d_array n α ≃ (Π i, α i) := -⟨d_array.read, d_array.mk, λ ⟨f⟩, rfl, λ f, rfl⟩ - -/-- The natural equivalence between length-`n` arrays and functions from `fin n`. -/ -def array_equiv_fin (n : ℕ) (α : Type*) : array n α ≃ (fin n → α) := -d_array_equiv_fin _ - -/-- The natural equivalence between length-`n` vectors and length-`n` arrays. -/ -def vector_equiv_array (α : Type*) (n : ℕ) : vector α n ≃ array n α := -(vector_equiv_fin _ _).trans (array_equiv_fin _ _).symm - -end equiv - -namespace array -open function -variable {n : ℕ} - -instance : traversable (array n) := -@equiv.traversable (flip vector n) _ (λ α, equiv.vector_equiv_array α n) _ - -instance : is_lawful_traversable (array n) := -@equiv.is_lawful_traversable (flip vector n) _ (λ α, equiv.vector_equiv_array α n) _ _ - -end array diff --git a/src/data/fintype/array.lean b/src/data/fintype/array.lean new file mode 100644 index 0000000000000..5211c8cebc957 --- /dev/null +++ b/src/data/fintype/array.lean @@ -0,0 +1,20 @@ +/- +Copyright (c) 2017 Mario Carneiro. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Mario Carneiro +-/ +import data.fintype.pi +import logic.equiv.array + +/-! +# `array n α` is a fintype when `α` is. +-/ + +variables {α : Type*} + +instance d_array.fintype {n : ℕ} {α : fin n → Type*} + [∀ n, fintype (α n)] : fintype (d_array n α) := +fintype.of_equiv _ (equiv.d_array_equiv_fin _).symm + +instance array.fintype {n : ℕ} {α : Type*} [fintype α] : fintype (array n α) := +d_array.fintype diff --git a/src/data/fintype/vector.lean b/src/data/fintype/vector.lean index 07ce3c1ef85c8..9b4b98cf30adb 100644 --- a/src/data/fintype/vector.lean +++ b/src/data/fintype/vector.lean @@ -4,22 +4,14 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ import data.fintype.pi -import data.array.lemmas import data.sym.basic /-! -# `vector α n` is a fintype when `α` is. +# `vector α n` and `sym α n` are fintypes when `α` is. -/ variables {α : Type*} -instance d_array.fintype {n : ℕ} {α : fin n → Type*} - [∀ n, fintype (α n)] : fintype (d_array n α) := -fintype.of_equiv _ (equiv.d_array_equiv_fin _).symm - -instance array.fintype {n : ℕ} {α : Type*} [fintype α] : fintype (array n α) := -d_array.fintype - instance vector.fintype [fintype α] {n : ℕ} : fintype (vector α n) := fintype.of_equiv _ (equiv.vector_equiv_fin _ _).symm diff --git a/src/logic/equiv/array.lean b/src/logic/equiv/array.lean new file mode 100644 index 0000000000000..5e1f98c7d71cf --- /dev/null +++ b/src/logic/equiv/array.lean @@ -0,0 +1,52 @@ +/- +Copyright (c) 2018 Mario Carneiro. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Mario Carneiro +-/ +import data.vector.basic +import logic.equiv.list +import control.traversable.equiv + +/-! +# Equivalences involving `array` + +We keep this separate from the file containing `list`-like equivalences as those have no future +in mathlib4. +-/ + +namespace equiv + +/-- The natural equivalence between length-`n` heterogeneous arrays +and dependent functions from `fin n`. -/ +def d_array_equiv_fin {n : ℕ} (α : fin n → Type*) : d_array n α ≃ (Π i, α i) := +⟨d_array.read, d_array.mk, λ ⟨f⟩, rfl, λ f, rfl⟩ + +/-- The natural equivalence between length-`n` arrays and functions from `fin n`. -/ +def array_equiv_fin (n : ℕ) (α : Type*) : array n α ≃ (fin n → α) := +d_array_equiv_fin _ + +/-- The natural equivalence between length-`n` vectors and length-`n` arrays. -/ +def vector_equiv_array (α : Type*) (n : ℕ) : vector α n ≃ array n α := +(vector_equiv_fin _ _).trans (array_equiv_fin _ _).symm + +end equiv + +namespace array +open function +variable {n : ℕ} + +instance : traversable (array n) := +@equiv.traversable (flip vector n) _ (λ α, equiv.vector_equiv_array α n) _ + +instance : is_lawful_traversable (array n) := +@equiv.is_lawful_traversable (flip vector n) _ (λ α, equiv.vector_equiv_array α n) _ _ + +end array + +/-- If `α` is encodable, then so is `array n α`. -/ +instance _root_.array.encodable {α} [encodable α] {n} : encodable (array n α) := +encodable.of_equiv _ (equiv.array_equiv_fin _ _) + +/-- If `α` is countable, then so is `array n α`. -/ +instance _root_.array.countable {α} [countable α] {n} : countable (array n α) := +countable.of_equiv _ (equiv.vector_equiv_array _ _) diff --git a/src/logic/equiv/list.lean b/src/logic/equiv/list.lean index 63cda4d0f5634..71a3d9e03d6fd 100644 --- a/src/logic/equiv/list.lean +++ b/src/logic/equiv/list.lean @@ -3,8 +3,8 @@ Copyright (c) 2018 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ -import data.array.lemmas import data.finset.sort +import data.vector.basic import logic.denumerable /-! @@ -130,14 +130,6 @@ of_equiv _ (equiv.vector_equiv_fin _ _).symm instance fin_pi (n) (π : fin n → Type*) [∀ i, encodable (π i)] : encodable (Π i, π i) := of_equiv _ (equiv.pi_equiv_subtype_sigma (fin n) π) -/-- If `α` is encodable, then so is `array n α`. -/ -instance _root_.array.encodable [encodable α] {n} : encodable (array n α) := -of_equiv _ (equiv.array_equiv_fin _ _) - -/-- If `α` is countable, then so is `array n α`. -/ -instance _root_.array.countable [countable α] {n} : countable (array n α) := -countable.of_equiv _ (equiv.vector_equiv_array _ _) - /-- If `α` is encodable, then so is `finset α`. -/ instance _root_.finset.encodable [encodable α] : encodable (finset α) := by haveI := decidable_eq_of_encodable α; exact From 4060545f1f2b865a399c46f14391e2dba0fe3667 Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Thu, 19 Jan 2023 18:26:19 +0000 Subject: [PATCH 0300/1291] feat(algebra/group/basic): reduce additivity checks to one case (#18080) Co-authored-by: Yury G. Kudryashov --- src/algebra/group/basic.lean | 44 ++++++++++++++++++++++++++++++++++++ 1 file changed, 44 insertions(+) diff --git a/src/algebra/group/basic.lean b/src/algebra/group/basic.lean index 6740b509dae66..9d06c6de2b904 100644 --- a/src/algebra/group/basic.lean +++ b/src/algebra/group/basic.lean @@ -651,3 +651,47 @@ lemma bit1_sub [has_one M] (a b : M) : bit1 (a - b) = bit1 a - bit0 b := (congr_arg (+ (1 : M)) $ bit0_sub a b : _).trans $ sub_add_eq_add_sub _ _ _ end subtraction_comm_monoid + +section multiplicative + +variables [monoid β] (p r : α → α → Prop) [is_total α r] (f : α → α → β) + +@[to_additive additive_of_symmetric_of_is_total] +lemma multiplicative_of_symmetric_of_is_total + (hsymm : symmetric p) (hf_swap : ∀ {a b}, p a b → f a b * f b a = 1) + (hmul : ∀ {a b c}, r a b → r b c → p a b → p b c → p a c → f a c = f a b * f b c) + {a b c : α} (pab : p a b) (pbc : p b c) (pac : p a c) : f a c = f a b * f b c := +begin + suffices : ∀ {b c}, r b c → p a b → p b c → p a c → f a c = f a b * f b c, + { obtain rbc | rcb := total_of r b c, + { exact this rbc pab pbc pac }, + { rw [this rcb pac (hsymm pbc) pab, mul_assoc, hf_swap (hsymm pbc), mul_one] } }, + intros b c rbc pab pbc pac, + obtain rab | rba := total_of r a b, + { exact hmul rab rbc pab pbc pac }, + rw [← one_mul (f a c), ← hf_swap pab, mul_assoc], + obtain rac | rca := total_of r a c, + { rw [hmul rba rac (hsymm pab) pac pbc] }, + { rw [hmul rbc rca pbc (hsymm pac) (hsymm pab), mul_assoc, hf_swap (hsymm pac), mul_one] }, +end + +/-- If a binary function from a type equipped with a total relation `r` to a monoid is + anti-symmetric (i.e. satisfies `f a b * f b a = 1`), in order to show it is multiplicative + (i.e. satisfies `f a c = f a b * f b c`), we may assume `r a b` and `r b c` are satisfied. + We allow restricting to a subset specified by a predicate `p`. -/ +@[to_additive additive_of_is_total "If a binary function from a type equipped with a total relation + `r` to an additive monoid is anti-symmetric (i.e. satisfies `f a b + f b a = 0`), in order to show + it is additive (i.e. satisfies `f a c = f a b + f b c`), we may assume `r a b` and `r b c` + are satisfied. We allow restricting to a subset specified by a predicate `p`."] +lemma multiplicative_of_is_total (p : α → Prop) + (hswap : ∀ {a b}, p a → p b → f a b * f b a = 1) + (hmul : ∀ {a b c}, r a b → r b c → p a → p b → p c → f a c = f a b * f b c) + {a b c : α} (pa : p a) (pb : p b) (pc : p c) : f a c = f a b * f b c := +begin + apply multiplicative_of_symmetric_of_is_total (λ a b, p a ∧ p b) r f (λ _ _, and.swap), + { simp_rw and_imp, exact @hswap }, + { exact λ a b c rab rbc pab pbc pac, hmul rab rbc pab.1 pab.2 pac.2 }, + exacts [⟨pa, pb⟩, ⟨pb, pc⟩, ⟨pa, pc⟩], +end + +end multiplicative From 1126441d6bccf98c81214a0780c73d499f6721fe Mon Sep 17 00:00:00 2001 From: Reid Barton Date: Thu, 19 Jan 2023 18:26:20 +0000 Subject: [PATCH 0301/1291] fix(topology/basic): docstring typo (#18229) --- src/topology/basic.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/topology/basic.lean b/src/topology/basic.lean index 101947cac8602..e3784ae39ea07 100644 --- a/src/topology/basic.lean +++ b/src/topology/basic.lean @@ -11,7 +11,7 @@ import order.filter.lift /-! # Basic theory of topological spaces. -The main definition is the type class `topological space α` which endows a type `α` with a topology. +The main definition is the type class `topological_space α` which endows a type `α` with a topology. Then `set α` gets predicates `is_open`, `is_closed` and functions `interior`, `closure` and `frontier`. Each point `x` of `α` gets a neighborhood filter `𝓝 x`. A filter `F` on `α` has `x` as a cluster point if `cluster_pt x F : 𝓝 x ⊓ F ≠ ⊥`. A map `f : ι → α` clusters at `x` From 0f6670b8af2dff699de1c0b4b49039b31bc13c46 Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Thu, 19 Jan 2023 21:39:18 +0000 Subject: [PATCH 0302/1291] chore(group_theory/subgroup/basic): split out some self-contained sections (#17862) This is a very big file by mathlib standards at over 3000 lines. It seems helpful to reduce that by splitting some of the less fundamental definitions into their own files. --- src/algebra/periodic.lean | 1 + src/group_theory/coset.lean | 1 + src/group_theory/group_action/conj_act.lean | 2 +- .../group_action/fixing_subgroup.lean | 2 +- src/group_theory/solvable.lean | 1 + .../specific_groups/alternating.lean | 1 + src/group_theory/specific_groups/cyclic.lean | 1 + src/group_theory/subgroup/actions.lean | 71 ++++ src/group_theory/subgroup/basic.lean | 347 ------------------ src/group_theory/subgroup/mul_opposite.lean | 54 +++ src/group_theory/subgroup/pointwise.lean | 2 +- src/group_theory/subgroup/saturated.lean | 58 +++ src/group_theory/subgroup/simple.lean | 70 ++++ src/group_theory/subgroup/zpowers.lean | 171 +++++++++ src/linear_algebra/ray.lean | 1 + 15 files changed, 433 insertions(+), 350 deletions(-) create mode 100644 src/group_theory/subgroup/actions.lean create mode 100644 src/group_theory/subgroup/mul_opposite.lean create mode 100644 src/group_theory/subgroup/saturated.lean create mode 100644 src/group_theory/subgroup/simple.lean create mode 100644 src/group_theory/subgroup/zpowers.lean diff --git a/src/algebra/periodic.lean b/src/algebra/periodic.lean index 442c4cafb585b..9404f8a061254 100644 --- a/src/algebra/periodic.lean +++ b/src/algebra/periodic.lean @@ -8,6 +8,7 @@ import algebra.module.basic import algebra.order.archimedean import data.int.parity import group_theory.coset +import group_theory.subgroup.zpowers /-! # Periodicity diff --git a/src/group_theory/coset.lean b/src/group_theory/coset.lean index 48d24dd1d9afa..2eee06e3571d5 100644 --- a/src/group_theory/coset.lean +++ b/src/group_theory/coset.lean @@ -6,6 +6,7 @@ Authors: Mitchell Rowett, Scott Morrison import algebra.quotient import group_theory.group_action.basic +import group_theory.subgroup.mul_opposite import tactic.group /-! diff --git a/src/group_theory/group_action/conj_act.lean b/src/group_theory/group_action/conj_act.lean index 6c216817766fe..5a9ed7a2ffb8c 100644 --- a/src/group_theory/group_action/conj_act.lean +++ b/src/group_theory/group_action/conj_act.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes -/ import group_theory.group_action.basic -import group_theory.subgroup.basic +import group_theory.subgroup.zpowers import algebra.group_ring_action.basic /-! # Conjugation action of a group on itself diff --git a/src/group_theory/group_action/fixing_subgroup.lean b/src/group_theory/group_action/fixing_subgroup.lean index 3718d4c8ac4da..ba7d078189e48 100644 --- a/src/group_theory/group_action/fixing_subgroup.lean +++ b/src/group_theory/group_action/fixing_subgroup.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Antoine Chambert-Loir -/ -import group_theory.subgroup.basic +import group_theory.subgroup.actions import group_theory.group_action.basic /-! diff --git a/src/group_theory/solvable.lean b/src/group_theory/solvable.lean index e9ae039668dda..d02f872a54c4d 100644 --- a/src/group_theory/solvable.lean +++ b/src/group_theory/solvable.lean @@ -7,6 +7,7 @@ Authors: Jordan Brown, Thomas Browning, Patrick Lutz import data.fin.vec_notation import group_theory.abelianization import group_theory.perm.via_embedding +import group_theory.subgroup.simple import set_theory.cardinal.basic /-! diff --git a/src/group_theory/specific_groups/alternating.lean b/src/group_theory/specific_groups/alternating.lean index 5fddc8b655f95..5f4a13d2e7245 100644 --- a/src/group_theory/specific_groups/alternating.lean +++ b/src/group_theory/specific_groups/alternating.lean @@ -6,6 +6,7 @@ Authors: Aaron Anderson import algebra.group.conj_finite import group_theory.perm.fin +import group_theory.subgroup.simple import tactic.interval_cases /-! diff --git a/src/group_theory/specific_groups/cyclic.lean b/src/group_theory/specific_groups/cyclic.lean index 21e961156bfdb..9a48297f65a54 100644 --- a/src/group_theory/specific_groups/cyclic.lean +++ b/src/group_theory/specific_groups/cyclic.lean @@ -7,6 +7,7 @@ Authors: Johannes Hölzl import algebra.big_operators.order import data.nat.totient import group_theory.order_of_element +import group_theory.subgroup.simple import tactic.group import group_theory.exponent diff --git a/src/group_theory/subgroup/actions.lean b/src/group_theory/subgroup/actions.lean new file mode 100644 index 0000000000000..c3d8d0ad3fa79 --- /dev/null +++ b/src/group_theory/subgroup/actions.lean @@ -0,0 +1,71 @@ +/- +Copyright (c) 2021 Eric Wieser. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Eric Wieser +-/ + +import group_theory.subgroup.basic + +/-! +# Actions by `subgroup`s + +These are just copies of the definitions about `submonoid` starting from `submonoid.mul_action`. + +## Tags +subgroup, subgroups + +-/ + +namespace subgroup + +variables {G : Type*} [group G] +variables {α β : Type*} + +/-- The action by a subgroup is the action by the underlying group. -/ +@[to_additive /-"The additive action by an add_subgroup is the action by the underlying +add_group. "-/] +instance [mul_action G α] (S : subgroup G) : mul_action S α := +S.to_submonoid.mul_action + +@[to_additive] +lemma smul_def [mul_action G α] {S : subgroup G} (g : S) (m : α) : g • m = (g : G) • m := rfl + +@[to_additive] +instance smul_comm_class_left + [mul_action G β] [has_smul α β] [smul_comm_class G α β] (S : subgroup G) : + smul_comm_class S α β := +S.to_submonoid.smul_comm_class_left + +@[to_additive] +instance smul_comm_class_right + [has_smul α β] [mul_action G β] [smul_comm_class α G β] (S : subgroup G) : + smul_comm_class α S β := +S.to_submonoid.smul_comm_class_right + +/-- Note that this provides `is_scalar_tower S G G` which is needed by `smul_mul_assoc`. -/ +instance + [has_smul α β] [mul_action G α] [mul_action G β] [is_scalar_tower G α β] (S : subgroup G) : + is_scalar_tower S α β := +S.to_submonoid.is_scalar_tower + +instance [mul_action G α] [has_faithful_smul G α] (S : subgroup G) : + has_faithful_smul S α := +S.to_submonoid.has_faithful_smul + +/-- The action by a subgroup is the action by the underlying group. -/ +instance [add_monoid α] [distrib_mul_action G α] (S : subgroup G) : distrib_mul_action S α := +S.to_submonoid.distrib_mul_action + +/-- The action by a subgroup is the action by the underlying group. -/ +instance [monoid α] [mul_distrib_mul_action G α] (S : subgroup G) : mul_distrib_mul_action S α := +S.to_submonoid.mul_distrib_mul_action + +/-- The center of a group acts commutatively on that group. -/ +instance center.smul_comm_class_left : smul_comm_class (center G) G G := +submonoid.center.smul_comm_class_left + +/-- The center of a group acts commutatively on that group. -/ +instance center.smul_comm_class_right : smul_comm_class G (center G) G := +submonoid.center.smul_comm_class_right + +end subgroup diff --git a/src/group_theory/subgroup/basic.lean b/src/group_theory/subgroup/basic.lean index a5c2c2982e21d..c84d3f099aae4 100644 --- a/src/group_theory/subgroup/basic.lean +++ b/src/group_theory/subgroup/basic.lean @@ -76,8 +76,6 @@ Definitions in the file: * `monoid_hom.eq_locus f g` : given group homomorphisms `f`, `g`, the elements of `G` such that `f x = g x` form a subgroup of `G` -* `is_simple_group G` : a class indicating that a group has exactly two normal subgroups - ## Implementation notes Subgroup inclusion is denoted `≤` rather than `⊆`, although `∈` is defined as @@ -2613,157 +2611,6 @@ hH.comap _ instance subgroup.normal_subgroup_of {H N : subgroup G} [N.normal] : (N.subgroup_of H).normal := subgroup.normal_comap _ -namespace subgroup - -/-- The subgroup generated by an element. -/ -def zpowers (g : G) : subgroup G := -subgroup.copy (zpowers_hom G g).range (set.range ((^) g : ℤ → G)) rfl - -@[simp] lemma mem_zpowers (g : G) : g ∈ zpowers g := ⟨1, zpow_one _⟩ - -lemma zpowers_eq_closure (g : G) : zpowers g = closure {g} := -by { ext, exact mem_closure_singleton.symm } - -@[simp] lemma range_zpowers_hom (g : G) : (zpowers_hom G g).range = zpowers g := rfl - -lemma zpowers_subset {a : G} {K : subgroup G} (h : a ∈ K) : zpowers a ≤ K := -λ x hx, match x, hx with _, ⟨i, rfl⟩ := K.zpow_mem h i end - -lemma mem_zpowers_iff {g h : G} : - h ∈ zpowers g ↔ ∃ (k : ℤ), g ^ k = h := -iff.rfl - -@[simp] lemma zpow_mem_zpowers (g : G) (k : ℤ) : g^k ∈ zpowers g := -mem_zpowers_iff.mpr ⟨k, rfl⟩ - -@[simp] lemma npow_mem_zpowers (g : G) (k : ℕ) : g^k ∈ zpowers g := -(zpow_coe_nat g k) ▸ zpow_mem_zpowers g k - -@[simp] lemma forall_zpowers {x : G} {p : zpowers x → Prop} : - (∀ g, p g) ↔ ∀ m : ℤ, p ⟨x ^ m, m, rfl⟩ := -set.forall_subtype_range_iff - -@[simp] lemma exists_zpowers {x : G} {p : zpowers x → Prop} : - (∃ g, p g) ↔ ∃ m : ℤ, p ⟨x ^ m, m, rfl⟩ := -set.exists_subtype_range_iff - -lemma forall_mem_zpowers {x : G} {p : G → Prop} : - (∀ g ∈ zpowers x, p g) ↔ ∀ m : ℤ, p (x ^ m) := -set.forall_range_iff - -lemma exists_mem_zpowers {x : G} {p : G → Prop} : - (∃ g ∈ zpowers x, p g) ↔ ∃ m : ℤ, p (x ^ m) := -set.exists_range_iff - -instance (a : G) : countable (zpowers a) := -((zpowers_hom G a).range_restrict_surjective.comp multiplicative.of_add.surjective).countable - -end subgroup - -namespace add_subgroup - -/-- The subgroup generated by an element. -/ -def zmultiples (a : A) : add_subgroup A := -add_subgroup.copy (zmultiples_hom A a).range (set.range ((• a) : ℤ → A)) rfl - -@[simp] lemma range_zmultiples_hom (a : A) : (zmultiples_hom A a).range = zmultiples a := rfl - -attribute [to_additive add_subgroup.zmultiples] subgroup.zpowers -attribute [to_additive add_subgroup.mem_zmultiples] subgroup.mem_zpowers -attribute [to_additive add_subgroup.zmultiples_eq_closure] subgroup.zpowers_eq_closure -attribute [to_additive add_subgroup.range_zmultiples_hom] subgroup.range_zpowers_hom -attribute [to_additive add_subgroup.zmultiples_subset] subgroup.zpowers_subset -attribute [to_additive add_subgroup.mem_zmultiples_iff] subgroup.mem_zpowers_iff -attribute [to_additive add_subgroup.zsmul_mem_zmultiples] subgroup.zpow_mem_zpowers -attribute [to_additive add_subgroup.nsmul_mem_zmultiples] subgroup.npow_mem_zpowers -attribute [to_additive add_subgroup.forall_zmultiples] subgroup.forall_zpowers -attribute [to_additive add_subgroup.forall_mem_zmultiples] subgroup.forall_mem_zpowers -attribute [to_additive add_subgroup.exists_zmultiples] subgroup.exists_zpowers -attribute [to_additive add_subgroup.exists_mem_zmultiples] subgroup.exists_mem_zpowers - -instance (a : A) : countable (zmultiples a) := -(zmultiples_hom A a).range_restrict_surjective.countable - -section ring - -variables {R : Type*} [ring R] (r : R) (k : ℤ) - -@[simp] lemma int_cast_mul_mem_zmultiples : - ↑(k : ℤ) * r ∈ zmultiples r := -by simpa only [← zsmul_eq_mul] using zsmul_mem_zmultiples r k - -@[simp] lemma int_cast_mem_zmultiples_one : - ↑(k : ℤ) ∈ zmultiples (1 : R) := -mem_zmultiples_iff.mp ⟨k, by simp⟩ - -end ring - -end add_subgroup - -@[simp, to_additive map_zmultiples] lemma monoid_hom.map_zpowers (f : G →* N) (x : G) : - (subgroup.zpowers x).map f = subgroup.zpowers (f x) := -by rw [subgroup.zpowers_eq_closure, subgroup.zpowers_eq_closure, f.map_closure, set.image_singleton] - -lemma int.mem_zmultiples_iff {a b : ℤ} : - b ∈ add_subgroup.zmultiples a ↔ a ∣ b := -exists_congr (λ k, by rw [mul_comm, eq_comm, ← smul_eq_mul]) - -lemma of_mul_image_zpowers_eq_zmultiples_of_mul { x : G } : - additive.of_mul '' ((subgroup.zpowers x) : set G) = add_subgroup.zmultiples (additive.of_mul x) := -begin - ext y, - split, - { rintro ⟨z, ⟨m, hm⟩, hz2⟩, - use m, - simp only, - rwa [← of_mul_zpow, hm] }, - { rintros ⟨n, hn⟩, - refine ⟨x ^ n, ⟨n, rfl⟩, _⟩, - rwa of_mul_zpow } -end - -lemma of_add_image_zmultiples_eq_zpowers_of_add {x : A} : - multiplicative.of_add '' ((add_subgroup.zmultiples x) : set A) = - subgroup.zpowers (multiplicative.of_add x) := -begin - symmetry, - rw equiv.eq_image_iff_symm_image_eq, - exact of_mul_image_zpowers_eq_zmultiples_of_mul, -end - -namespace subgroup - -@[to_additive zmultiples_is_commutative] -instance zpowers_is_commutative (g : G) : (zpowers g).is_commutative := -⟨⟨λ ⟨_, _, h₁⟩ ⟨_, _, h₂⟩, by rw [subtype.ext_iff, coe_mul, coe_mul, - subtype.coe_mk, subtype.coe_mk, ←h₁, ←h₂, zpow_mul_comm]⟩⟩ - -@[simp, to_additive zmultiples_le] -lemma zpowers_le {g : G} {H : subgroup G} : zpowers g ≤ H ↔ g ∈ H := -by rw [zpowers_eq_closure, closure_le, set.singleton_subset_iff, set_like.mem_coe] - -@[simp, to_additive zmultiples_eq_bot] lemma zpowers_eq_bot {g : G} : zpowers g = ⊥ ↔ g = 1 := -by rw [eq_bot_iff, zpowers_le, mem_bot] - -@[simp, to_additive zmultiples_zero_eq_bot] lemma zpowers_one_eq_bot : - subgroup.zpowers (1 : G) = ⊥ := -subgroup.zpowers_eq_bot.mpr rfl - -@[to_additive] lemma centralizer_closure (S : set G) : - (closure S).centralizer = ⨅ g ∈ S, (zpowers g).centralizer := -le_antisymm (le_infi $ λ g, le_infi $ λ hg, centralizer_le $ zpowers_le.2 $ subset_closure hg) - $ le_centralizer_iff.1 $ (closure_le _).2 - $ λ g, set_like.mem_coe.2 ∘ zpowers_le.1 ∘ le_centralizer_iff.1 ∘ infi_le_of_le g ∘ infi_le _ - -@[to_additive] lemma center_eq_infi (S : set G) (hS : closure S = ⊤) : - center G = ⨅ g ∈ S, centralizer (zpowers g) := -by rw [←centralizer_top, ←hS, centralizer_closure] - -@[to_additive] lemma center_eq_infi' (S : set G) (hS : closure S = ⊤) : - center G = ⨅ g : S, centralizer (zpowers g) := -by rw [center_eq_infi S hS, ←infi_subtype''] - -end subgroup namespace monoid_hom @@ -2857,49 +2704,6 @@ end⟩ end subgroup -section -variables (G) (A) - -/-- A `group` is simple when it has exactly two normal `subgroup`s. -/ -class is_simple_group extends nontrivial G : Prop := -(eq_bot_or_eq_top_of_normal : ∀ H : subgroup G, H.normal → H = ⊥ ∨ H = ⊤) - -/-- An `add_group` is simple when it has exactly two normal `add_subgroup`s. -/ -class is_simple_add_group extends nontrivial A : Prop := -(eq_bot_or_eq_top_of_normal : ∀ H : add_subgroup A, H.normal → H = ⊥ ∨ H = ⊤) - -attribute [to_additive] is_simple_group - -variables {G} {A} - -@[to_additive] -lemma subgroup.normal.eq_bot_or_eq_top [is_simple_group G] {H : subgroup G} (Hn : H.normal) : - H = ⊥ ∨ H = ⊤ := -is_simple_group.eq_bot_or_eq_top_of_normal H Hn - -namespace is_simple_group - -@[to_additive] -instance {C : Type*} [comm_group C] [is_simple_group C] : - is_simple_order (subgroup C) := -⟨λ H, H.normal_of_comm.eq_bot_or_eq_top⟩ - -open _root_.subgroup - -@[to_additive] -lemma is_simple_group_of_surjective {H : Type*} [group H] [is_simple_group G] - [nontrivial H] (f : G →* H) (hf : function.surjective f) : - is_simple_group H := -⟨nontrivial.exists_pair_ne, λ H iH, begin - refine ((iH.comap f).eq_bot_or_eq_top).imp (λ h, _) (λ h, _), - { rw [←map_bot f, ←h, map_comap_eq_self_of_surjective hf] }, - { rw [←comap_top f] at h, exact comap_injective hf h } -end⟩ - -end is_simple_group - -end - namespace subgroup section subgroup_normal @@ -3071,154 +2875,3 @@ begin end end is_conj - -/-! ### Actions by `subgroup`s - -These are just copies of the definitions about `submonoid` starting from `submonoid.mul_action`. --/ -section actions - -namespace subgroup - -variables {α β : Type*} - -/-- The action by a subgroup is the action by the underlying group. -/ -@[to_additive /-"The additive action by an add_subgroup is the action by the underlying -add_group. "-/] -instance [mul_action G α] (S : subgroup G) : mul_action S α := -S.to_submonoid.mul_action - -@[to_additive] -lemma smul_def [mul_action G α] {S : subgroup G} (g : S) (m : α) : g • m = (g : G) • m := rfl - -@[to_additive] -instance smul_comm_class_left - [mul_action G β] [has_smul α β] [smul_comm_class G α β] (S : subgroup G) : - smul_comm_class S α β := -S.to_submonoid.smul_comm_class_left - -@[to_additive] -instance smul_comm_class_right - [has_smul α β] [mul_action G β] [smul_comm_class α G β] (S : subgroup G) : - smul_comm_class α S β := -S.to_submonoid.smul_comm_class_right - -/-- Note that this provides `is_scalar_tower S G G` which is needed by `smul_mul_assoc`. -/ -instance - [has_smul α β] [mul_action G α] [mul_action G β] [is_scalar_tower G α β] (S : subgroup G) : - is_scalar_tower S α β := -S.to_submonoid.is_scalar_tower - -instance [mul_action G α] [has_faithful_smul G α] (S : subgroup G) : - has_faithful_smul S α := -S.to_submonoid.has_faithful_smul - -/-- The action by a subgroup is the action by the underlying group. -/ -instance [add_monoid α] [distrib_mul_action G α] (S : subgroup G) : distrib_mul_action S α := -S.to_submonoid.distrib_mul_action - -/-- The action by a subgroup is the action by the underlying group. -/ -instance [monoid α] [mul_distrib_mul_action G α] (S : subgroup G) : mul_distrib_mul_action S α := -S.to_submonoid.mul_distrib_mul_action - -/-- The center of a group acts commutatively on that group. -/ -instance center.smul_comm_class_left : smul_comm_class (center G) G G := -submonoid.center.smul_comm_class_left - -/-- The center of a group acts commutatively on that group. -/ -instance center.smul_comm_class_right : smul_comm_class G (center G) G := -submonoid.center.smul_comm_class_right - -end subgroup - -end actions - -/-! ### Mul-opposite subgroups -/ - -section mul_opposite - -namespace subgroup - -/-- A subgroup `H` of `G` determines a subgroup `H.opposite` of the opposite group `Gᵐᵒᵖ`. -/ -@[to_additive "An additive subgroup `H` of `G` determines an additive subgroup `H.opposite` of the - opposite additive group `Gᵃᵒᵖ`."] -def opposite : subgroup G ≃ subgroup Gᵐᵒᵖ := -{ to_fun := λ H, { carrier := mul_opposite.unop ⁻¹' (H : set G), - one_mem' := H.one_mem, - mul_mem' := λ a b ha hb, H.mul_mem hb ha, - inv_mem' := λ a, H.inv_mem }, - inv_fun := λ H, { carrier := mul_opposite.op ⁻¹' (H : set Gᵐᵒᵖ), - one_mem' := H.one_mem, - mul_mem' := λ a b ha hb, H.mul_mem hb ha, - inv_mem' := λ a, H.inv_mem }, - left_inv := λ H, set_like.coe_injective rfl, - right_inv := λ H, set_like.coe_injective rfl } - -/-- Bijection between a subgroup `H` and its opposite. -/ -@[to_additive "Bijection between an additive subgroup `H` and its opposite.", simps] -def opposite_equiv (H : subgroup G) : H ≃ H.opposite := -mul_opposite.op_equiv.subtype_equiv $ λ _, iff.rfl - -@[to_additive] instance (H : subgroup G) [encodable H] : encodable H.opposite := -encodable.of_equiv H H.opposite_equiv.symm - -@[to_additive] instance (H : subgroup G) [countable H] : countable H.opposite := -countable.of_equiv H H.opposite_equiv - -@[to_additive] lemma smul_opposite_mul {H : subgroup G} (x g : G) (h : H.opposite) : - h • (g * x) = g * (h • x) := -begin - cases h, - simp [(•), mul_assoc], -end - -end subgroup - -end mul_opposite - -/-! ### Saturated subgroups -/ - -section saturated - -namespace subgroup - -/-- A subgroup `H` of `G` is *saturated* if for all `n : ℕ` and `g : G` with `g^n ∈ H` -we have `n = 0` or `g ∈ H`. -/ -@[to_additive "An additive subgroup `H` of `G` is *saturated* if -for all `n : ℕ` and `g : G` with `n•g ∈ H` we have `n = 0` or `g ∈ H`."] -def saturated (H : subgroup G) : Prop := ∀ ⦃n g⦄, g ^ n ∈ H → n = 0 ∨ g ∈ H - -@[to_additive] lemma saturated_iff_npow {H : subgroup G} : - saturated H ↔ (∀ (n : ℕ) (g : G), g ^ n ∈ H → n = 0 ∨ g ∈ H) := iff.rfl - -@[to_additive] lemma saturated_iff_zpow {H : subgroup G} : - saturated H ↔ (∀ (n : ℤ) (g : G), g ^ n ∈ H → n = 0 ∨ g ∈ H) := -begin - split, - { rintros hH ⟨n⟩ g hgn, - { simp only [int.coe_nat_eq_zero, int.of_nat_eq_coe, zpow_coe_nat] at hgn ⊢, - exact hH hgn }, - { suffices : g ^ (n+1) ∈ H, - { refine (hH this).imp _ id, simp only [is_empty.forall_iff, nat.succ_ne_zero], }, - simpa only [inv_mem_iff, zpow_neg_succ_of_nat] using hgn, } }, - { intros h n g hgn, - specialize h n g, - simp only [int.coe_nat_eq_zero, zpow_coe_nat] at h, - apply h hgn } -end - -end subgroup - -namespace add_subgroup - -lemma ker_saturated {A₁ A₂ : Type*} [add_comm_group A₁] [add_comm_group A₂] - [no_zero_smul_divisors ℕ A₂] (f : A₁ →+ A₂) : - (f.ker).saturated := -begin - intros n g hg, - simpa only [f.mem_ker, nsmul_eq_smul, f.map_nsmul, smul_eq_zero] using hg -end - -end add_subgroup - -end saturated diff --git a/src/group_theory/subgroup/mul_opposite.lean b/src/group_theory/subgroup/mul_opposite.lean new file mode 100644 index 0000000000000..93f590238c684 --- /dev/null +++ b/src/group_theory/subgroup/mul_opposite.lean @@ -0,0 +1,54 @@ +/- +Copyright (c) 2022 Alex Kontorovich. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Alex Kontorovich +-/ + +import group_theory.subgroup.actions + +/-! +# Mul-opposite subgroups + +## Tags +subgroup, subgroups + +-/ + +variables {G : Type*} [group G] + +namespace subgroup + +/-- A subgroup `H` of `G` determines a subgroup `H.opposite` of the opposite group `Gᵐᵒᵖ`. -/ +@[to_additive "An additive subgroup `H` of `G` determines an additive subgroup `H.opposite` of the + opposite additive group `Gᵃᵒᵖ`."] +def opposite : subgroup G ≃ subgroup Gᵐᵒᵖ := +{ to_fun := λ H, { carrier := mul_opposite.unop ⁻¹' (H : set G), + one_mem' := H.one_mem, + mul_mem' := λ a b ha hb, H.mul_mem hb ha, + inv_mem' := λ a, H.inv_mem }, + inv_fun := λ H, { carrier := mul_opposite.op ⁻¹' (H : set Gᵐᵒᵖ), + one_mem' := H.one_mem, + mul_mem' := λ a b ha hb, H.mul_mem hb ha, + inv_mem' := λ a, H.inv_mem }, + left_inv := λ H, set_like.coe_injective rfl, + right_inv := λ H, set_like.coe_injective rfl } + +/-- Bijection between a subgroup `H` and its opposite. -/ +@[to_additive "Bijection between an additive subgroup `H` and its opposite.", simps] +def opposite_equiv (H : subgroup G) : H ≃ H.opposite := +mul_opposite.op_equiv.subtype_equiv $ λ _, iff.rfl + +@[to_additive] instance (H : subgroup G) [encodable H] : encodable H.opposite := +encodable.of_equiv H H.opposite_equiv.symm + +@[to_additive] instance (H : subgroup G) [countable H] : countable H.opposite := +countable.of_equiv H H.opposite_equiv + +@[to_additive] lemma smul_opposite_mul {H : subgroup G} (x g : G) (h : H.opposite) : + h • (g * x) = g * (h • x) := +begin + cases h, + simp [(•), mul_assoc], +end + +end subgroup diff --git a/src/group_theory/subgroup/pointwise.lean b/src/group_theory/subgroup/pointwise.lean index 67d21096238a2..3a565dfb0313f 100644 --- a/src/group_theory/subgroup/pointwise.lean +++ b/src/group_theory/subgroup/pointwise.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Eric Wieser. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Eric Wieser -/ -import group_theory.subgroup.basic +import group_theory.subgroup.mul_opposite import group_theory.submonoid.pointwise import group_theory.group_action.conj_act diff --git a/src/group_theory/subgroup/saturated.lean b/src/group_theory/subgroup/saturated.lean new file mode 100644 index 0000000000000..e22b4921d7165 --- /dev/null +++ b/src/group_theory/subgroup/saturated.lean @@ -0,0 +1,58 @@ +/- +Copyright (c) 2021 Johan Commelin. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johan Commelin +-/ + +import group_theory.subgroup.basic + +/-! +# Saturated subgroups + +## Tags +subgroup, subgroups + +-/ + +namespace subgroup + +variables {G : Type*} [group G] + +/-- A subgroup `H` of `G` is *saturated* if for all `n : ℕ` and `g : G` with `g^n ∈ H` +we have `n = 0` or `g ∈ H`. -/ +@[to_additive "An additive subgroup `H` of `G` is *saturated* if +for all `n : ℕ` and `g : G` with `n•g ∈ H` we have `n = 0` or `g ∈ H`."] +def saturated (H : subgroup G) : Prop := ∀ ⦃n g⦄, g ^ n ∈ H → n = 0 ∨ g ∈ H + +@[to_additive] lemma saturated_iff_npow {H : subgroup G} : + saturated H ↔ (∀ (n : ℕ) (g : G), g ^ n ∈ H → n = 0 ∨ g ∈ H) := iff.rfl + +@[to_additive] lemma saturated_iff_zpow {H : subgroup G} : + saturated H ↔ (∀ (n : ℤ) (g : G), g ^ n ∈ H → n = 0 ∨ g ∈ H) := +begin + split, + { rintros hH ⟨n⟩ g hgn, + { simp only [int.coe_nat_eq_zero, int.of_nat_eq_coe, zpow_coe_nat] at hgn ⊢, + exact hH hgn }, + { suffices : g ^ (n+1) ∈ H, + { refine (hH this).imp _ id, simp only [is_empty.forall_iff, nat.succ_ne_zero], }, + simpa only [inv_mem_iff, zpow_neg_succ_of_nat] using hgn, } }, + { intros h n g hgn, + specialize h n g, + simp only [int.coe_nat_eq_zero, zpow_coe_nat] at h, + apply h hgn } +end + +end subgroup + +namespace add_subgroup + +lemma ker_saturated {A₁ A₂ : Type*} [add_comm_group A₁] [add_comm_group A₂] + [no_zero_smul_divisors ℕ A₂] (f : A₁ →+ A₂) : + (f.ker).saturated := +begin + intros n g hg, + simpa only [f.mem_ker, nsmul_eq_smul, f.map_nsmul, smul_eq_zero] using hg +end + +end add_subgroup diff --git a/src/group_theory/subgroup/simple.lean b/src/group_theory/subgroup/simple.lean new file mode 100644 index 0000000000000..f043081c9206e --- /dev/null +++ b/src/group_theory/subgroup/simple.lean @@ -0,0 +1,70 @@ +/- +Copyright (c) 2021 Aaron Anderson. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Aaron Anderson +-/ + +import group_theory.subgroup.actions + +/-! +# Simple groups + +This file defines `is_simple_group G`, a class indicating that a group has exactly two normal +subgroups. + +## Main definitions + +- `is_simple_group G`, a class indicating that a group has exactly two normal subgroups. + +## Tags +subgroup, subgroups + +-/ + +set_option old_structure_cmd true + +variables {G : Type*} [group G] +variables {A : Type*} [add_group A] + +section +variables (G) (A) + +/-- A `group` is simple when it has exactly two normal `subgroup`s. -/ +class is_simple_group extends nontrivial G : Prop := +(eq_bot_or_eq_top_of_normal : ∀ H : subgroup G, H.normal → H = ⊥ ∨ H = ⊤) + +/-- An `add_group` is simple when it has exactly two normal `add_subgroup`s. -/ +class is_simple_add_group extends nontrivial A : Prop := +(eq_bot_or_eq_top_of_normal : ∀ H : add_subgroup A, H.normal → H = ⊥ ∨ H = ⊤) + +attribute [to_additive] is_simple_group + +variables {G} {A} + +@[to_additive] +lemma subgroup.normal.eq_bot_or_eq_top [is_simple_group G] {H : subgroup G} (Hn : H.normal) : + H = ⊥ ∨ H = ⊤ := +is_simple_group.eq_bot_or_eq_top_of_normal H Hn + +namespace is_simple_group + +@[to_additive] +instance {C : Type*} [comm_group C] [is_simple_group C] : + is_simple_order (subgroup C) := +⟨λ H, H.normal_of_comm.eq_bot_or_eq_top⟩ + +open _root_.subgroup + +@[to_additive] +lemma is_simple_group_of_surjective {H : Type*} [group H] [is_simple_group G] + [nontrivial H] (f : G →* H) (hf : function.surjective f) : + is_simple_group H := +⟨nontrivial.exists_pair_ne, λ H iH, begin + refine ((iH.comap f).eq_bot_or_eq_top).imp (λ h, _) (λ h, _), + { rw [←map_bot f, ←h, map_comap_eq_self_of_surjective hf] }, + { rw [←comap_top f] at h, exact comap_injective hf h } +end⟩ + +end is_simple_group + +end diff --git a/src/group_theory/subgroup/zpowers.lean b/src/group_theory/subgroup/zpowers.lean new file mode 100644 index 0000000000000..b7a89fb456edd --- /dev/null +++ b/src/group_theory/subgroup/zpowers.lean @@ -0,0 +1,171 @@ +/- +Copyright (c) 2020 Chris Hughes. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Chris Hughes +-/ + +import group_theory.subgroup.basic + +/-! +# Subgroups generated by an element + +## Tags +subgroup, subgroups + +-/ + +variables {G : Type*} [group G] +variables {A : Type*} [add_group A] +variables {N : Type*} [group N] + +namespace subgroup + +/-- The subgroup generated by an element. -/ +def zpowers (g : G) : subgroup G := +subgroup.copy (zpowers_hom G g).range (set.range ((^) g : ℤ → G)) rfl + +@[simp] lemma mem_zpowers (g : G) : g ∈ zpowers g := ⟨1, zpow_one _⟩ + +lemma zpowers_eq_closure (g : G) : zpowers g = closure {g} := +by { ext, exact mem_closure_singleton.symm } + +@[simp] lemma range_zpowers_hom (g : G) : (zpowers_hom G g).range = zpowers g := rfl + +lemma zpowers_subset {a : G} {K : subgroup G} (h : a ∈ K) : zpowers a ≤ K := +λ x hx, match x, hx with _, ⟨i, rfl⟩ := K.zpow_mem h i end + +lemma mem_zpowers_iff {g h : G} : + h ∈ zpowers g ↔ ∃ (k : ℤ), g ^ k = h := +iff.rfl + +@[simp] lemma zpow_mem_zpowers (g : G) (k : ℤ) : g^k ∈ zpowers g := +mem_zpowers_iff.mpr ⟨k, rfl⟩ + +@[simp] lemma npow_mem_zpowers (g : G) (k : ℕ) : g^k ∈ zpowers g := +(zpow_coe_nat g k) ▸ zpow_mem_zpowers g k + +@[simp] lemma forall_zpowers {x : G} {p : zpowers x → Prop} : + (∀ g, p g) ↔ ∀ m : ℤ, p ⟨x ^ m, m, rfl⟩ := +set.forall_subtype_range_iff + +@[simp] lemma exists_zpowers {x : G} {p : zpowers x → Prop} : + (∃ g, p g) ↔ ∃ m : ℤ, p ⟨x ^ m, m, rfl⟩ := +set.exists_subtype_range_iff + +lemma forall_mem_zpowers {x : G} {p : G → Prop} : + (∀ g ∈ zpowers x, p g) ↔ ∀ m : ℤ, p (x ^ m) := +set.forall_range_iff + +lemma exists_mem_zpowers {x : G} {p : G → Prop} : + (∃ g ∈ zpowers x, p g) ↔ ∃ m : ℤ, p (x ^ m) := +set.exists_range_iff + +instance (a : G) : countable (zpowers a) := +((zpowers_hom G a).range_restrict_surjective.comp multiplicative.of_add.surjective).countable + +end subgroup + +namespace add_subgroup + +/-- The subgroup generated by an element. -/ +def zmultiples (a : A) : add_subgroup A := +add_subgroup.copy (zmultiples_hom A a).range (set.range ((• a) : ℤ → A)) rfl + +@[simp] lemma range_zmultiples_hom (a : A) : (zmultiples_hom A a).range = zmultiples a := rfl + +attribute [to_additive add_subgroup.zmultiples] subgroup.zpowers +attribute [to_additive add_subgroup.mem_zmultiples] subgroup.mem_zpowers +attribute [to_additive add_subgroup.zmultiples_eq_closure] subgroup.zpowers_eq_closure +attribute [to_additive add_subgroup.range_zmultiples_hom] subgroup.range_zpowers_hom +attribute [to_additive add_subgroup.zmultiples_subset] subgroup.zpowers_subset +attribute [to_additive add_subgroup.mem_zmultiples_iff] subgroup.mem_zpowers_iff +attribute [to_additive add_subgroup.zsmul_mem_zmultiples] subgroup.zpow_mem_zpowers +attribute [to_additive add_subgroup.nsmul_mem_zmultiples] subgroup.npow_mem_zpowers +attribute [to_additive add_subgroup.forall_zmultiples] subgroup.forall_zpowers +attribute [to_additive add_subgroup.forall_mem_zmultiples] subgroup.forall_mem_zpowers +attribute [to_additive add_subgroup.exists_zmultiples] subgroup.exists_zpowers +attribute [to_additive add_subgroup.exists_mem_zmultiples] subgroup.exists_mem_zpowers + +instance (a : A) : countable (zmultiples a) := +(zmultiples_hom A a).range_restrict_surjective.countable + +section ring + +variables {R : Type*} [ring R] (r : R) (k : ℤ) + +@[simp] lemma int_cast_mul_mem_zmultiples : + ↑(k : ℤ) * r ∈ zmultiples r := +by simpa only [← zsmul_eq_mul] using zsmul_mem_zmultiples r k + +@[simp] lemma int_cast_mem_zmultiples_one : + ↑(k : ℤ) ∈ zmultiples (1 : R) := +mem_zmultiples_iff.mp ⟨k, by simp⟩ + +end ring + +end add_subgroup + +@[simp, to_additive map_zmultiples] lemma monoid_hom.map_zpowers (f : G →* N) (x : G) : + (subgroup.zpowers x).map f = subgroup.zpowers (f x) := +by rw [subgroup.zpowers_eq_closure, subgroup.zpowers_eq_closure, f.map_closure, set.image_singleton] + +lemma int.mem_zmultiples_iff {a b : ℤ} : + b ∈ add_subgroup.zmultiples a ↔ a ∣ b := +exists_congr (λ k, by rw [mul_comm, eq_comm, ← smul_eq_mul]) + +lemma of_mul_image_zpowers_eq_zmultiples_of_mul { x : G } : + additive.of_mul '' ((subgroup.zpowers x) : set G) = add_subgroup.zmultiples (additive.of_mul x) := +begin + ext y, + split, + { rintro ⟨z, ⟨m, hm⟩, hz2⟩, + use m, + simp only, + rwa [← of_mul_zpow, hm] }, + { rintros ⟨n, hn⟩, + refine ⟨x ^ n, ⟨n, rfl⟩, _⟩, + rwa of_mul_zpow } +end + +lemma of_add_image_zmultiples_eq_zpowers_of_add {x : A} : + multiplicative.of_add '' ((add_subgroup.zmultiples x) : set A) = + subgroup.zpowers (multiplicative.of_add x) := +begin + symmetry, + rw equiv.eq_image_iff_symm_image_eq, + exact of_mul_image_zpowers_eq_zmultiples_of_mul, +end + +namespace subgroup + +@[to_additive zmultiples_is_commutative] +instance zpowers_is_commutative (g : G) : (zpowers g).is_commutative := +⟨⟨λ ⟨_, _, h₁⟩ ⟨_, _, h₂⟩, by rw [subtype.ext_iff, coe_mul, coe_mul, + subtype.coe_mk, subtype.coe_mk, ←h₁, ←h₂, zpow_mul_comm]⟩⟩ + +@[simp, to_additive zmultiples_le] +lemma zpowers_le {g : G} {H : subgroup G} : zpowers g ≤ H ↔ g ∈ H := +by rw [zpowers_eq_closure, closure_le, set.singleton_subset_iff, set_like.mem_coe] + +@[simp, to_additive zmultiples_eq_bot] lemma zpowers_eq_bot {g : G} : zpowers g = ⊥ ↔ g = 1 := +by rw [eq_bot_iff, zpowers_le, mem_bot] + +@[simp, to_additive zmultiples_zero_eq_bot] lemma zpowers_one_eq_bot : + subgroup.zpowers (1 : G) = ⊥ := +subgroup.zpowers_eq_bot.mpr rfl + +@[to_additive] lemma centralizer_closure (S : set G) : + (closure S).centralizer = ⨅ g ∈ S, (zpowers g).centralizer := +le_antisymm (le_infi $ λ g, le_infi $ λ hg, centralizer_le $ zpowers_le.2 $ subset_closure hg) + $ le_centralizer_iff.1 $ (closure_le _).2 + $ λ g, set_like.mem_coe.2 ∘ zpowers_le.1 ∘ le_centralizer_iff.1 ∘ infi_le_of_le g ∘ infi_le _ + +@[to_additive] lemma center_eq_infi (S : set G) (hS : closure S = ⊤) : + center G = ⨅ g ∈ S, centralizer (zpowers g) := +by rw [←centralizer_top, ←hS, centralizer_closure] + +@[to_additive] lemma center_eq_infi' (S : set G) (hS : closure S = ⊤) : + center G = ⨅ g : S, centralizer (zpowers g) := +by rw [center_eq_infi S hS, ←infi_subtype''] + +end subgroup diff --git a/src/linear_algebra/ray.lean b/src/linear_algebra/ray.lean index 9d9a2801a262f..c12aec99af056 100644 --- a/src/linear_algebra/ray.lean +++ b/src/linear_algebra/ray.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Joseph Myers. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Joseph Myers -/ +import group_theory.subgroup.actions import linear_algebra.linear_independent /-! From f808feb6c18afddb25e66a71d317643cf7fb5fbb Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Fri, 20 Jan 2023 00:18:26 +0000 Subject: [PATCH 0303/1291] feat(data/list/alist): recursion on `alist` using `insert` (#15434) --- src/data/list/alist.lean | 49 ++++++++++++++++++++++++++++++++++++++++ src/data/list/sigma.lean | 6 +++++ 2 files changed, 55 insertions(+) diff --git a/src/data/list/alist.lean b/src/data/list/alist.lean index 8648c1c2b183e..79b6778ecb5a8 100644 --- a/src/data/list/alist.lean +++ b/src/data/list/alist.lean @@ -205,6 +205,11 @@ theorem insert_entries_of_neg {a} {b : β a} {s : alist β} (h : a ∉ s) : (insert a b s).entries = ⟨a, b⟩ :: s.entries := by rw [insert_entries, kerase_of_not_mem_keys h] +-- Todo: rename to `insert_of_not_mem`. +theorem insert_of_neg {a} {b : β a} {s : alist β} (h : a ∉ s) : + insert a b s = ⟨⟨a, b⟩ :: s.entries, nodupkeys_cons.2 ⟨h, s.2⟩⟩ := +ext $ insert_entries_of_neg h + @[simp] theorem insert_empty (a) (b : β a) : insert a b ∅ = singleton a b := rfl @[simp] theorem mem_insert {a a'} {b' : β a'} (s : alist β) : @@ -250,6 +255,50 @@ ext $ by simp only [alist.insert_entries, list.kerase_cons_eq, and_self, alist.s theorem to_alist_cons (a : α) (b : β a) (xs : list (sigma β)) : list.to_alist (⟨a,b⟩ :: xs) = insert a b xs.to_alist := rfl +theorem mk_cons_eq_insert (c : sigma β) (l : list (sigma β)) (h : (c :: l).nodupkeys) : + (⟨c :: l, h⟩ : alist β) = insert c.1 c.2 ⟨l, nodupkeys_of_nodupkeys_cons h⟩ := +by simpa [insert] using (kerase_of_not_mem_keys $ not_mem_keys_of_nodupkeys_cons h).symm + +/-- Recursion on an `alist`, using `insert`. Use as `induction l using alist.insert_rec`. -/ +@[elab_as_eliminator] def insert_rec {C : alist β → Sort*} (H0 : C ∅) + (IH : Π (a : α) (b : β a) (l : alist β) (h : a ∉ l), C l → C (l.insert a b)) : Π l : alist β, C l +| ⟨[], _⟩ := H0 +| ⟨c :: l, h⟩ := begin + rw mk_cons_eq_insert, + refine IH _ _ _ _ (insert_rec _), + exact not_mem_keys_of_nodupkeys_cons h +end + +-- Test that the `induction` tactic works on `insert_rec`. +example (l : alist β) : true := by induction l using alist.insert_rec; trivial + +@[simp] theorem insert_rec_empty {C : alist β → Sort*} (H0 : C ∅) + (IH : Π (a : α) (b : β a) (l : alist β) (h : a ∉ l), C l → C (l.insert a b)) : + @insert_rec α β _ C H0 IH ∅ = H0 := +by { change @insert_rec α β _ C H0 IH ⟨[], _⟩ = H0, rw insert_rec } + +theorem insert_rec_insert {C : alist β → Sort*} (H0 : C ∅) + (IH : Π (a : α) (b : β a) (l : alist β) (h : a ∉ l), C l → C (l.insert a b)) + {c : sigma β} {l : alist β} (h : c.1 ∉ l) : + @insert_rec α β _ C H0 IH (l.insert c.1 c.2) = IH c.1 c.2 l h (@insert_rec α β _ C H0 IH l) := +begin + cases l with l hl, + suffices : @insert_rec α β _ C H0 IH ⟨c :: l, nodupkeys_cons.2 ⟨h, hl⟩⟩ == + IH c.1 c.2 ⟨l, hl⟩ h (@insert_rec α β _ C H0 IH ⟨l, hl⟩), + { cases c, + apply eq_of_heq, + convert this; + rw insert_of_neg h }, + rw insert_rec, + apply cast_heq +end + +theorem recursion_insert_mk {C : alist β → Sort*} (H0 : C ∅) + (IH : Π (a : α) (b : β a) (l : alist β) (h : a ∉ l), C l → C (l.insert a b)) + {a : α} (b : β a) {l : alist β} (h : a ∉ l) : + @insert_rec α β _ C H0 IH (l.insert a b) = IH a b l h (@insert_rec α β _ C H0 IH l) := +@insert_rec_insert α β _ C H0 IH ⟨a, b⟩ l h + /-! ### extract -/ /-- Erase a key from the map, and return the corresponding value, if found. -/ diff --git a/src/data/list/sigma.lean b/src/data/list/sigma.lean index 10f58ae9d0eca..b24f8564ce41a 100644 --- a/src/data/list/sigma.lean +++ b/src/data/list/sigma.lean @@ -80,6 +80,12 @@ nodupkeys_iff_pairwise.1 h nodupkeys (s::l) ↔ s.1 ∉ l.keys ∧ nodupkeys l := by simp [keys, nodupkeys] +theorem not_mem_keys_of_nodupkeys_cons {s : sigma β} {l : list (sigma β)} (h : nodupkeys (s :: l)) : + s.1 ∉ l.keys := (nodupkeys_cons.1 h).1 + +theorem nodupkeys_of_nodupkeys_cons {s : sigma β} {l : list (sigma β)} (h : nodupkeys (s :: l)) : + nodupkeys l := (nodupkeys_cons.1 h).2 + theorem nodupkeys.eq_of_fst_eq {l : list (sigma β)} (nd : nodupkeys l) {s s' : sigma β} (h : s ∈ l) (h' : s' ∈ l) : s.1 = s'.1 → s = s' := From 3f0c8b329fd917175c304db864297c8c36ac2ba2 Mon Sep 17 00:00:00 2001 From: Moritz Firsching Date: Fri, 20 Jan 2023 04:49:06 +0000 Subject: [PATCH 0304/1291] chore(*): detect blobs in port_status (#18237) see https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/Port.20Status.20Webpage Co-authored-by: Moritz Firsching --- scripts/port_status.py | 5 +++++ 1 file changed, 5 insertions(+) diff --git a/scripts/port_status.py b/scripts/port_status.py index d4456453058a1..42766e5c9809b 100755 --- a/scripts/port_status.py +++ b/scripts/port_status.py @@ -71,6 +71,11 @@ def mk_label(path: Path) -> str: for node in graph.nodes: if data[node].mathlib3_hash: verified[node] = data[node].mathlib3_hash + find_blobs_command = ['git', 'cat-file', '-t', data[node].mathlib3_hash] + hash_type = subprocess.check_output(find_blobs_command) + # the hash_type should be commits mostly, we are not interested in blobs + if b'blob\n' == hash_type: + break git_command = ['git', 'diff', '--quiet', f'--ignore-matching-lines={comment_git_re}', data[node].mathlib3_hash + "..HEAD", "--", "src" + os.sep + node.replace('.', os.sep) + ".lean"] From 52e2fbbd9b519e6981e4b53aedb26a2d52548d3a Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Fri, 20 Jan 2023 07:10:13 +0000 Subject: [PATCH 0305/1291] feat(special_functions/gamma): Bohr-Mollerup theorem (#18188) Add the characterisation of the Gamma function as the unique log-convex function with the correct functional equation. --- src/analysis/special_functions/gamma.lean | 458 +++++++++++++++++- src/analysis/special_functions/gaussian.lean | 2 +- src/analysis/special_functions/log/basic.lean | 27 ++ src/data/real/conjugate_exponents.lean | 4 + 4 files changed, 466 insertions(+), 25 deletions(-) diff --git a/src/analysis/special_functions/gamma.lean b/src/analysis/special_functions/gamma.lean index e0d175e767c5b..d6775ee17112c 100644 --- a/src/analysis/special_functions/gamma.lean +++ b/src/analysis/special_functions/gamma.lean @@ -17,7 +17,33 @@ integral `Γ(s) = ∫ x in Ioi 0, exp (-x) * x ^ (s - 1)` in the range where thi We show that this integral satisfies `Γ(1) = 1` and `Γ(s + 1) = s * Γ(s)`; hence we can define `Γ(s)` for all `s` as the unique function satisfying this recurrence and agreeing with Euler's integral in the convergence range. In the complex case we also prove that the resulting function is -holomorphic on `ℂ` away from the points `{-n : n ∈ ℤ}`. +holomorphic on `ℂ` away from the points `{-n : n ∈ ℕ}`. + +## Main statements (real case) + +* `real.Gamma` : the `Γ` function (of a real variable). +* `real.Gamma_eq_integral` : for `0 < s`, `Γ(s)` agrees with Euler's integral + `∫ (x:ℝ) in Ioi 0, exp (-x) * x ^ (s - 1)` +* `real.Gamma_add_one` : for all `s : ℝ` with `s ≠ 0`, we have `Γ(s + 1) = s Γ(s)`. +* `real.Gamma_nat_eq_factorial` : for all `n : ℕ` we have `Γ (n + 1) = n!`. +* `real.differentiable_at_Gamma` : `Γ` is real-differentiable at all `s : ℝ` with + `s ∉ {-n : n ∈ ℕ}`. +* `real.Gamma_ne_zero`: for all `s : ℝ` with `s ∉ {-n : n ∈ ℕ}` we have `Γ s ≠ 0`. +* `real.tendsto_log_Gamma`: for all `0 < s`, the limit as `n → ∞` of the sequence + `n ↦ s log n + log n! - (log s + ... + log (s + n))` is `log Γ(s)`. +* `real.convex_on_log_Gamma` : `log ∘ Γ` is convex on `Ioi 0`. +* `real.eq_Gamma_of_log_convex` : the Bohr-Mollerup theorem, which states that the `Γ` function is + the unique log-convex, positive-valued function on `Ioi 0` satisfying the functional equation + and having `Γ 1 = 1`. + +## Main statements (complex case) + +* `complex.Gamma` : the `Γ` function (of a complex variable). +* `complex.Gamma_eq_integral` : for `0 < re s`, `Γ(s)` agrees with Euler's integral. +* `complex.Gamma_add_one` : for all `s : ℂ` with `s ≠ 0`, we have `Γ(s + 1) = s Γ(s)`. +* `complex.Gamma_nat_eq_factorial` : for all `n : ℕ` we have `Γ (n + 1) = n!`. +* `complex.differentiable_at_Gamma` : `Γ` is complex-differentiable at all `s : ℂ` with + `s ∉ {-n : n ∈ ℕ}`. ## Tags @@ -26,7 +52,7 @@ Gamma noncomputable theory open filter interval_integral set real measure_theory asymptotics -open_locale topological_space +open_locale nat topological_space ennreal big_operators lemma integral_exp_neg_Ioi : ∫ (x : ℝ) in Ioi 0, exp (-x) = 1 := begin @@ -52,13 +78,7 @@ begin exact (tendsto_exp_mul_div_rpow_at_top s (1 / 2) one_half_pos).inv_tendsto_at_top, end -/-- Euler's integral for the `Γ` function (of a real variable `s`), defined as -`∫ x in Ioi 0, exp (-x) * x ^ (s - 1)`. - -See `Gamma_integral_convergent` for a proof of the convergence of the integral for `0 < s`. -/ -def Gamma_integral (s : ℝ) : ℝ := ∫ x in Ioi (0:ℝ), exp (-x) * x ^ (s - 1) - -/-- The integral defining the `Γ` function converges for positive real `s`. -/ +/-- The Euler integral for the `Γ` function converges for positive real `s`. -/ lemma Gamma_integral_convergent {s : ℝ} (h : 0 < s) : integrable_on (λ x:ℝ, exp (-x) * x ^ (s - 1)) (Ioi 0) := begin @@ -74,9 +94,6 @@ begin exact or.inl ((zero_lt_one : (0 : ℝ) < 1).trans_le hx).ne' } end -lemma Gamma_integral_one : Gamma_integral 1 = 1 := -by simpa only [Gamma_integral, sub_self, rpow_zero, mul_one] using integral_exp_neg_Ioi - end real namespace complex @@ -116,9 +133,9 @@ See `complex.Gamma_integral_convergent` for a proof of the convergence of the in def Gamma_integral (s : ℂ) : ℂ := ∫ x in Ioi (0:ℝ), ↑(-x).exp * ↑x ^ (s - 1) lemma Gamma_integral_of_real (s : ℝ) : - Gamma_integral ↑s = ↑(s.Gamma_integral) := + Gamma_integral ↑s = ↑(∫ x:ℝ in Ioi 0, real.exp (-x) * x ^ (s - 1)) := begin - rw [real.Gamma_integral, ←_root_.integral_of_real], + rw [Gamma_integral, ←_root_.integral_of_real], refine set_integral_congr measurable_set_Ioi _, intros x hx, dsimp only, rw [of_real_mul, of_real_cpow (mem_Ioi.mp hx).le], @@ -126,10 +143,8 @@ begin end lemma Gamma_integral_one : Gamma_integral 1 = 1 := -begin - rw [←of_real_one, Gamma_integral_of_real, of_real_inj], - exact real.Gamma_integral_one, -end +by simpa only [←of_real_one, Gamma_integral_of_real, of_real_inj, sub_self, + rpow_zero, mul_one] using integral_exp_neg_Ioi end complex @@ -327,16 +342,18 @@ begin field_simp, ring, end -theorem Gamma_eq_integral (s : ℂ) (hs : 0 < s.re) : Gamma s = Gamma_integral s := +theorem Gamma_eq_integral {s : ℂ} (hs : 0 < s.re) : Gamma s = Gamma_integral s := Gamma_eq_Gamma_aux s 0 (by { norm_cast, linarith }) -theorem Gamma_nat_eq_factorial (n : ℕ) : Gamma (n+1) = nat.factorial n := +lemma Gamma_one : Gamma 1 = 1 := +by { rw Gamma_eq_integral, simpa using Gamma_integral_one, simp } + +theorem Gamma_nat_eq_factorial (n : ℕ) : Gamma (n+1) = n! := begin induction n with n hn, - { rw [nat.cast_zero, zero_add], rw Gamma_eq_integral, - simpa using Gamma_integral_one, simp,}, - rw (Gamma_add_one n.succ $ nat.cast_ne_zero.mpr $ nat.succ_ne_zero n), - { simp only [nat.cast_succ, nat.factorial_succ, nat.cast_mul], congr, exact hn }, + { simpa using Gamma_one }, + { rw (Gamma_add_one n.succ $ nat.cast_ne_zero.mpr $ nat.succ_ne_zero n), + simp only [nat.cast_succ, nat.factorial_succ, nat.cast_mul], congr, exact hn }, end end Gamma_def @@ -519,3 +536,396 @@ end end complex end Gamma_has_deriv + +namespace real + +/-- The `Γ` function (of a real variable `s`). -/ +@[pp_nodot] def Gamma (s : ℝ) : ℝ := (complex.Gamma s).re + +lemma Gamma_eq_integral {s : ℝ} (hs : 0 < s) : Gamma s = ∫ x in Ioi 0, exp (-x) * x ^ (s - 1) := +begin + rw [Gamma, complex.Gamma_eq_integral (by rwa complex.of_real_re : 0 < complex.re s)], + dsimp only [complex.Gamma_integral], + simp_rw [←complex.of_real_one, ←complex.of_real_sub], + suffices : ∫ (x : ℝ) in Ioi 0, ↑(exp (-x)) * (x : ℂ) ^ ((s - 1 : ℝ) : ℂ) = + ∫ (x : ℝ) in Ioi 0, ((exp (-x) * x ^ (s - 1) : ℝ) : ℂ), + { rw [this, _root_.integral_of_real, complex.of_real_re], }, + refine set_integral_congr measurable_set_Ioi (λ x hx, _), + push_cast, + rw complex.of_real_cpow (le_of_lt hx), + push_cast, +end + +lemma Gamma_add_one {s : ℝ} (hs : s ≠ 0) : Gamma (s + 1) = s * Gamma s := +begin + simp_rw Gamma, + rw [complex.of_real_add, complex.of_real_one, complex.Gamma_add_one, complex.of_real_mul_re], + rwa complex.of_real_ne_zero, +end + +lemma Gamma_one : Gamma 1 = 1 := +by rw [Gamma, complex.of_real_one, complex.Gamma_one, complex.one_re] + +theorem Gamma_nat_eq_factorial (n : ℕ) : Gamma (n + 1) = n! := +by rw [Gamma, complex.of_real_add, complex.of_real_nat_cast, complex.of_real_one, + complex.Gamma_nat_eq_factorial, ←complex.of_real_nat_cast, complex.of_real_re] + +lemma Gamma_pos_of_pos {s : ℝ} (hs : 0 < s) : 0 < Gamma s := +begin + rw Gamma_eq_integral hs, + have : function.support (λ (x : ℝ), exp (-x) * x ^ (s - 1)) ∩ Ioi 0 = Ioi 0, + { rw inter_eq_right_iff_subset, + intros x hx, + rw function.mem_support, + exact mul_ne_zero (exp_pos _).ne' (rpow_pos_of_pos hx _).ne' }, + rw set_integral_pos_iff_support_of_nonneg_ae, + { rw [this, volume_Ioi, ←ennreal.of_real_zero], + exact ennreal.of_real_lt_top }, + { refine eventually_of_mem (self_mem_ae_restrict measurable_set_Ioi) _, + exact λ x hx, (mul_pos (exp_pos _) (rpow_pos_of_pos hx _)).le }, + { exact Gamma_integral_convergent hs }, +end + +lemma Gamma_ne_zero {s : ℝ} (hs : ∀ m:ℕ, s + m ≠ 0) : Gamma s ≠ 0 := +begin + suffices : ∀ {n : ℕ}, (-(n:ℝ) < s) → Gamma s ≠ 0, + { apply this, + swap, use (⌊-s⌋₊ + 1), + rw [neg_lt, nat.cast_add, nat.cast_one], + exact nat.lt_floor_add_one _ }, + intro n, + induction n generalizing s, + { intro hs, + refine (Gamma_pos_of_pos _).ne', + rwa [nat.cast_zero, neg_zero] at hs }, + { intro hs', + have : Gamma (s + 1) ≠ 0, + { apply n_ih, + { intro m, + convert hs (1 + m) using 1, + push_cast, + ring, }, + { rw [nat.succ_eq_add_one, nat.cast_add, nat.cast_one, neg_add] at hs', + linarith } }, + rw [Gamma_add_one, mul_ne_zero_iff] at this, + { exact this.2 }, + { simpa using hs 0 } }, +end + +lemma differentiable_at_Gamma {s : ℝ} (hs : ∀ m:ℕ, s + m ≠ 0) : differentiable_at ℝ Gamma s := +begin + apply has_deriv_at.differentiable_at, + apply has_deriv_at.real_of_complex, + apply differentiable_at.has_deriv_at, + apply complex.differentiable_at_Gamma, + simp_rw [←complex.of_real_nat_cast, ←complex.of_real_add, complex.of_real_ne_zero], + exact hs, +end + +/-- Log-convexity of the Gamma function on the positive reals (stated in multiplicative form), +proved using the Hölder inequality applied to Euler's integral. -/ +lemma Gamma_mul_add_mul_le_rpow_Gamma_mul_rpow_Gamma {s t a b : ℝ} + (hs : 0 < s) (ht : 0 < t) (ha : 0 < a) (hb : 0 < b) (hab : a + b = 1) : + Gamma (a * s + b * t) ≤ Gamma s ^ a * Gamma t ^ b := +begin + -- We will apply Hölder's inequality, for the conjugate exponents `p = 1 / a` + -- and `q = 1 / b`, to the functions `f a s` and `f b t`, where `f` is as follows: + let f : ℝ → ℝ → ℝ → ℝ := λ c u x, exp (-c * x) * x ^ (c * (u - 1)), + have e : is_conjugate_exponent (1 / a) (1 / b) := real.is_conjugate_exponent_one_div ha hb hab, + have hab' : b = 1 - a := by linarith, + have hst : 0 < a * s + b * t := add_pos (mul_pos ha hs) (mul_pos hb ht), + -- some properties of f: + have posf : ∀ (c u x : ℝ), x ∈ Ioi (0:ℝ) → 0 ≤ f c u x := + λ c u x hx, mul_nonneg (exp_pos _).le (rpow_pos_of_pos hx _).le, + have posf' : ∀ (c u : ℝ), ∀ᵐ (x : ℝ) ∂volume.restrict (Ioi 0), 0 ≤ f c u x := + λ c u, (ae_restrict_iff' measurable_set_Ioi).mpr (ae_of_all _ (posf c u)), + have fpow : ∀ {c x : ℝ} (hc : 0 < c) (u : ℝ) (hx : 0 < x), + exp (-x) * x ^ (u - 1) = f c u x ^ (1 / c), + { intros c x hc u hx, + dsimp only [f], + rw [mul_rpow (exp_pos _).le ((rpow_nonneg_of_nonneg hx.le) _), ←exp_mul, ←rpow_mul hx.le], + congr' 2; + { field_simp [hc.ne'], ring } }, + -- show `f c u` is in `ℒp` for `p = 1/c`: + have f_mem_Lp : ∀ {c u : ℝ} (hc : 0 < c) (hu : 0 < u), + mem_ℒp (f c u) (ennreal.of_real (1 / c)) (volume.restrict (Ioi 0)), + { intros c u hc hu, + have A : ennreal.of_real (1 / c) ≠ 0, + by rwa [ne.def, ennreal.of_real_eq_zero, not_le, one_div_pos], + have B : ennreal.of_real (1 / c) ≠ ∞, from ennreal.of_real_ne_top, + rw [←mem_ℒp_norm_rpow_iff _ A B, ennreal.to_real_of_real (one_div_nonneg.mpr hc.le), + ennreal.div_self A B, mem_ℒp_one_iff_integrable], + { apply integrable.congr (Gamma_integral_convergent hu), + refine eventually_eq_of_mem (self_mem_ae_restrict measurable_set_Ioi) (λ x hx, _), + dsimp only, + rw fpow hc u hx, + congr' 1, + exact (norm_of_nonneg (posf _ _ x hx)).symm }, + { refine continuous_on.ae_strongly_measurable _ measurable_set_Ioi, + refine (continuous.continuous_on _).mul (continuous_at.continuous_on (λ x hx, _)), + { exact continuous_exp.comp (continuous_const.mul continuous_id'), }, + { exact continuous_at_rpow_const _ _ (or.inl (ne_of_lt hx).symm), } } }, + -- now apply Hölder: + rw [Gamma_eq_integral hs, Gamma_eq_integral ht, Gamma_eq_integral hst], + convert measure_theory.integral_mul_le_Lp_mul_Lq_of_nonneg e (posf' a s) (posf' b t) + (f_mem_Lp ha hs) (f_mem_Lp hb ht) using 1, + { refine set_integral_congr measurable_set_Ioi (λ x hx, _), + dsimp only [f], + have A : exp (-x) = exp (-a * x) * exp (-b * x), + { rw [←exp_add, ←add_mul, ←neg_add, hab, neg_one_mul] }, + have B : x ^ (a * s + b * t - 1) = (x ^ (a * (s - 1))) * (x ^ (b * (t - 1))), + { rw [←rpow_add hx, hab'], congr' 1, ring }, + rw [A, B], + ring }, + { rw [one_div_one_div, one_div_one_div], + congr' 2; + exact set_integral_congr measurable_set_Ioi (λ x hx, fpow (by assumption) _ hx) }, +end + +lemma convex_on_log_Gamma : convex_on ℝ (Ioi 0) (log ∘ Gamma) := +begin + refine convex_on_iff_forall_pos.mpr ⟨convex_Ioi _, λ x hx y hy a b ha hb hab, _⟩, + have : b = 1 - a := by linarith, subst this, + simp_rw [function.comp_app, smul_eq_mul], + rw [←log_rpow (Gamma_pos_of_pos hy), ←log_rpow (Gamma_pos_of_pos hx), + ←log_mul + ((rpow_pos_of_pos (Gamma_pos_of_pos hx) _).ne') (rpow_pos_of_pos (Gamma_pos_of_pos hy) _).ne', + log_le_log + (Gamma_pos_of_pos (add_pos (mul_pos ha hx) (mul_pos hb hy))) + (mul_pos + (rpow_pos_of_pos (Gamma_pos_of_pos hx) _) (rpow_pos_of_pos (Gamma_pos_of_pos hy) _))], + exact Gamma_mul_add_mul_le_rpow_Gamma_mul_rpow_Gamma hx hy ha hb hab, +end + +section bohr_mollerup + +/-! ## The Euler limit formula and the Bohr-Mollerup theorem + +In this section we prove two interrelated statements about the `Γ` function on the positive reals: + +* the Euler limit formula `real.tendsto_log_gamma_seq`, stating that the sequence + `x * log n + log n! - ∑ (m : ℕ) in finset.range (n + 1), log (x + m)` + tends to `log Γ(x)` as `n → ∞`. +* the Bohr-Mollerup theorem (`real.eq_Gamma_of_log_convex`) which states that `Γ` is the unique + *log-convex*, positive-real-valued function on the positive reals satisfying + `f (x + 1) = x f x` and `f 1 = 1`. + +To do this, we prove that any function satisfying the hypotheses of the Bohr--Mollerup theorem must +agree with the limit in the Gauss formula, so there is at most one such function. Then we show that +`Γ` satisfies these conditions. +-/ + +/-- The function `n ↦ x log n + log n! - (log x + ... + log (x + n))`, which we will show tends to +`log (Gamma x)` as `n → ∞`. -/ +def log_gamma_seq (x : ℝ) (n : ℕ) : ℝ := +x * log n + log n! - ∑ (m : ℕ) in finset.range (n + 1), log (x + m) + +/-! The following are auxiliary lemmas for the Bohr-Mollerup theorem, which are +placed in a separate namespace `bohr_mollerup` to avoid clutter. -/ +namespace bohr_mollerup + +variables {f : ℝ → ℝ} {x : ℝ} {n : ℕ} + +lemma f_nat_eq (hf_feq : ∀ {y:ℝ}, 0 < y → f (y + 1) = f y + log y) (hn : n ≠ 0) : + f n = f 1 + log (n - 1)! := +begin + refine nat.le_induction (by simp) (λ m hm IH, _) n (nat.one_le_iff_ne_zero.2 hn), + have A : 0 < (m : ℝ), from nat.cast_pos.2 hm, + simp only [hf_feq A, nat.cast_add, algebra_map.coe_one, nat.add_succ_sub_one, add_zero], + rw [IH, add_assoc, ← log_mul (nat.cast_ne_zero.mpr (nat.factorial_ne_zero _)) A.ne', + ← nat.cast_mul], + conv_rhs { rw [← nat.succ_pred_eq_of_pos hm, nat.factorial_succ, mul_comm] }, + congr, + exact (nat.succ_pred_eq_of_pos hm).symm +end + +lemma f_add_nat_eq (hf_feq : ∀ {y:ℝ}, 0 < y → f (y + 1) = f y + log y) (hx : 0 < x) (n : ℕ) : + f (x + n) = f x + ∑ (m : ℕ) in finset.range n, log (x + m) := +begin + induction n with n hn, + { simp }, + { have : x + n.succ = (x + n) + 1, + { push_cast, ring }, + rw [this, hf_feq, hn], + rw [finset.range_succ, finset.sum_insert (finset.not_mem_range_self)], + abel, + linarith [(nat.cast_nonneg n : 0 ≤ (n:ℝ))] }, +end + +/-- Linear upper bound for `f (x + n)` on unit interval -/ +lemma f_add_nat_le + (hf_conv : convex_on ℝ (Ioi 0) f) (hf_feq : ∀ {y:ℝ}, 0 < y → f (y + 1) = f y + log y) + (hn : n ≠ 0) (hx : 0 < x) (hx' : x ≤ 1) : + f (n + x) ≤ f n + x * log n := +begin + have hn': 0 < (n:ℝ) := nat.cast_pos.mpr (nat.pos_of_ne_zero hn), + have : f n + x * log n = (1 - x) * f n + x * f (n + 1), + { rw [hf_feq hn'], ring, }, + rw [this, (by ring : (n:ℝ) + x = (1 - x) * n + x * (n + 1))], + simpa only [smul_eq_mul] using hf_conv.2 hn' (by linarith : 0 < (n + 1 : ℝ)) + (by linarith : 0 ≤ 1 - x) hx.le (by linarith), +end + +/-- Linear lower bound for `f (x + n)` on unit interval -/ +lemma f_add_nat_ge + (hf_conv : convex_on ℝ (Ioi 0) f) (hf_feq : ∀ {y:ℝ}, 0 < y → f (y + 1) = f y + log y) + (hn : 2 ≤ n) (hx : 0 < x) : + f n + x * log (n - 1) ≤ f (n + x) := +begin + have npos : 0 < (n:ℝ) - 1, + { rw [←nat.cast_one, sub_pos, nat.cast_lt], linarith, }, + have c := (convex_on_iff_slope_mono_adjacent.mp $ hf_conv).2 + npos (by linarith : 0 < (n:ℝ) + x) (by linarith : (n:ℝ) - 1 < (n:ℝ)) (by linarith), + rw [add_sub_cancel', sub_sub_cancel, div_one] at c, + have : f (↑n - 1) = f n - log (↑n - 1), + { nth_rewrite_rhs 0 (by ring : (n:ℝ) = (↑n - 1) + 1), + rw [hf_feq npos, add_sub_cancel] }, + rwa [this, le_div_iff hx, sub_sub_cancel, le_sub_iff_add_le, mul_comm _ x, add_comm] at c, +end + +lemma log_gamma_seq_add_one (x : ℝ) (n : ℕ) : + log_gamma_seq (x + 1) n = log_gamma_seq x (n + 1) + log x - (x + 1) * (log (n + 1) - log n) := +begin + dsimp only [nat.factorial_succ, log_gamma_seq], + conv_rhs { rw [finset.sum_range_succ', nat.cast_zero, add_zero], }, + rw [nat.cast_mul, log_mul], rotate, + { rw nat.cast_ne_zero, exact nat.succ_ne_zero n }, + { rw nat.cast_ne_zero, exact nat.factorial_ne_zero n, }, + have : ∑ (m : ℕ) in finset.range (n + 1), log (x + 1 + ↑m) = + ∑ (k : ℕ) in finset.range (n + 1), log (x + ↑(k + 1)), + { refine finset.sum_congr (by refl) (λ m hm, _), + congr' 1, + push_cast, + abel }, + rw [←this, nat.cast_add_one n], + ring, +end + +lemma le_log_gamma_seq + (hf_conv : convex_on ℝ (Ioi 0) f) (hf_feq : ∀ {y:ℝ}, 0 < y → f (y + 1) = f y + log y) + (hx : 0 < x) (hx' : x ≤ 1) (n : ℕ) : + f x ≤ f 1 + x * log (n + 1) - x * log n + log_gamma_seq x n := +begin + rw [log_gamma_seq, ←add_sub_assoc, le_sub_iff_add_le, ←f_add_nat_eq @hf_feq hx, add_comm x], + refine (f_add_nat_le hf_conv @hf_feq (nat.add_one_ne_zero n) hx hx').trans (le_of_eq _), + rw [f_nat_eq @hf_feq (by linarith : n + 1 ≠ 0), nat.add_sub_cancel, nat.cast_add_one], + ring, +end + +lemma ge_log_gamma_seq + (hf_conv : convex_on ℝ (Ioi 0) f) (hf_feq : ∀ {y:ℝ}, 0 < y → f (y + 1) = f y + log y) + (hx : 0 < x) (hn : n ≠ 0) : + f 1 + log_gamma_seq x n ≤ f x := +begin + dsimp [log_gamma_seq], + rw [←add_sub_assoc, sub_le_iff_le_add, ←f_add_nat_eq @hf_feq hx, add_comm x _], + refine le_trans (le_of_eq _) (f_add_nat_ge hf_conv @hf_feq _ hx), + { rw [f_nat_eq @hf_feq, nat.add_sub_cancel, nat.cast_add_one, add_sub_cancel], + { ring }, + { exact nat.succ_ne_zero _} }, + { apply nat.succ_le_succ, + linarith [nat.pos_of_ne_zero hn] }, +end + +lemma tendsto_log_gamma_seq_of_le_one + (hf_conv : convex_on ℝ (Ioi 0) f) (hf_feq : ∀ {y:ℝ}, 0 < y → f (y + 1) = f y + log y) + (hx : 0 < x) (hx' : x ≤ 1) : + tendsto (log_gamma_seq x) at_top (𝓝 $ f x - f 1) := +begin + refine tendsto_of_tendsto_of_tendsto_of_le_of_le' _ tendsto_const_nhds _ _, + show ∀ᶠ (n : ℕ) in at_top, log_gamma_seq x n ≤ f x - f 1, + { refine eventually.mp (eventually_ne_at_top 0) (eventually_of_forall (λ n hn, _)), + exact le_sub_iff_add_le'.mpr (ge_log_gamma_seq hf_conv @hf_feq hx hn) }, + show ∀ᶠ (n : ℕ) in at_top, f x - f 1 - x * (log (n + 1) - log n) ≤ log_gamma_seq x n, + { refine eventually_of_forall (λ n, _), + rw [sub_le_iff_le_add', sub_le_iff_le_add'], + convert le_log_gamma_seq hf_conv @hf_feq hx hx' n using 1, + ring }, + { have : f x - f 1 = (f x - f 1) - x * 0 := by ring, + nth_rewrite 0 this, + exact tendsto.sub tendsto_const_nhds (tendsto_log_nat_add_one_sub_log.const_mul _), } +end + +lemma tendsto_log_gamma_seq + (hf_conv : convex_on ℝ (Ioi 0) f) (hf_feq : ∀ {y:ℝ}, 0 < y → f (y + 1) = f y + log y) + (hx : 0 < x) : + tendsto (log_gamma_seq x) at_top (𝓝 $ f x - f 1) := +begin + suffices : ∀ (m : ℕ), ↑m < x → x ≤ m + 1 → + tendsto (log_gamma_seq x) at_top (𝓝 $ f x - f 1), + { refine this (⌈x - 1⌉₊) _ _, + { rcases lt_or_le x 1, + { rwa [nat.ceil_eq_zero.mpr (by linarith : x - 1 ≤ 0), nat.cast_zero] }, + { convert nat.ceil_lt_add_one (by linarith : 0 ≤ x - 1), + abel } }, + { rw ←sub_le_iff_le_add, exact nat.le_ceil _}, }, + intro m, + induction m with m hm generalizing x, + { rw [nat.cast_zero, zero_add], + exact λ _ hx', tendsto_log_gamma_seq_of_le_one hf_conv @hf_feq hx hx' }, + { intros hy hy', + rw [nat.cast_succ, ←sub_le_iff_le_add] at hy', + rw [nat.cast_succ, ←lt_sub_iff_add_lt] at hy, + specialize hm ((nat.cast_nonneg _).trans_lt hy) hy hy', + -- now massage gauss_product n (x - 1) into gauss_product (n - 1) x + have : ∀ᶠ (n:ℕ) in at_top, log_gamma_seq (x - 1) n = log_gamma_seq x (n - 1) + + x * (log (↑(n - 1) + 1) - log ↑(n - 1)) - log (x - 1), + { refine eventually.mp (eventually_ge_at_top 1) (eventually_of_forall (λ n hn, _)), + have := log_gamma_seq_add_one (x - 1) (n - 1), + rw [sub_add_cancel, nat.sub_add_cancel hn] at this, + rw this, + ring }, + replace hm := ((tendsto.congr' this hm).add + (tendsto_const_nhds : tendsto (λ _, log (x - 1)) _ _)).comp (tendsto_add_at_top_nat 1), + have : + (λ (x_1 : ℕ), (λ (n : ℕ), log_gamma_seq x (n - 1) + + x * (log (↑(n - 1) + 1) - log ↑(n - 1)) - log (x - 1)) x_1 + + (λ (b : ℕ), log (x - 1)) x_1) ∘ (λ (a : ℕ), a + 1) = + λ n, log_gamma_seq x n + x * (log (↑n + 1) - log ↑n), + { ext1 n, + dsimp only [function.comp_app], + rw [sub_add_cancel, nat.add_sub_cancel] }, + rw this at hm, + convert hm.sub (tendsto_log_nat_add_one_sub_log.const_mul x) using 2, + { ext1 n, ring }, + { have := hf_feq ((nat.cast_nonneg m).trans_lt hy), + rw sub_add_cancel at this, + rw this, + ring } }, +end + +end bohr_mollerup + +lemma tendsto_log_Gamma {x : ℝ} (hx : 0 < x) : + tendsto (log_gamma_seq x) at_top (𝓝 $ log (Gamma x)) := +begin + have : log (Gamma x) = (log ∘ Gamma) x - (log ∘ Gamma) 1, + { simp_rw [function.comp_app, Gamma_one, log_one, sub_zero] }, + rw this, + refine bohr_mollerup.tendsto_log_gamma_seq convex_on_log_Gamma (λ y hy, _) hx, + rw [function.comp_app, Gamma_add_one hy.ne', log_mul hy.ne' (Gamma_pos_of_pos hy).ne', add_comm], +end + +/-- The **Bohr-Mollerup theorem**: the Gamma function is the *unique* log-convex, positive-valued +function on the positive reals which satisfies `f 1 = 1` and `f (x + 1) = x * f x` for all `x`. -/ +lemma eq_Gamma_of_log_convex {f : ℝ → ℝ} + (hf_conv : convex_on ℝ (Ioi 0) (log ∘ f)) + (hf_feq : ∀ {y:ℝ}, 0 < y → f (y + 1) = y * f y) + (hf_pos : ∀ {y:ℝ}, 0 < y → 0 < f y) + (hf_one : f 1 = 1) : + eq_on f Gamma (Ioi (0:ℝ)) := +begin + suffices : eq_on (log ∘ f) (log ∘ Gamma) (Ioi (0:ℝ)), + from λ x hx, log_inj_on_pos (hf_pos hx) (Gamma_pos_of_pos hx) (this hx), + intros x hx, + have e1 := bohr_mollerup.tendsto_log_gamma_seq hf_conv _ hx, + { rw [function.comp_app log f 1, hf_one, log_one, sub_zero] at e1, + exact tendsto_nhds_unique e1 (tendsto_log_Gamma hx) }, + { intros y hy, + rw [function.comp_app, hf_feq hy, log_mul hy.ne' (hf_pos hy).ne'], + ring } +end + +end bohr_mollerup + +end real diff --git a/src/analysis/special_functions/gaussian.lean b/src/analysis/special_functions/gaussian.lean index ef4792ec70e2d..8053436bc7b1a 100644 --- a/src/analysis/special_functions/gaussian.lean +++ b/src/analysis/special_functions/gaussian.lean @@ -320,7 +320,7 @@ begin have hh2 : (1 / 2 : ℂ).re = 1 / 2, { convert of_real_re (1 / 2 : ℝ) }, replace hh2 : 0 < (1 / 2 : ℂ).re := by { rw hh2, exact one_half_pos, }, - rw [Gamma_eq_integral _ hh2, hh, Gamma_integral_of_real, of_real_inj, real.Gamma_integral], + rw [Gamma_eq_integral hh2, hh, Gamma_integral_of_real, of_real_inj], -- now do change-of-variables rw ←integral_comp_rpow_Ioi_of_pos zero_lt_two, have : eq_on (λ x:ℝ, (2 * x^((2:ℝ) - 1)) • (real.exp (-x^(2:ℝ)) * (x^(2:ℝ)) ^ (1 / (2:ℝ) - 1))) diff --git a/src/analysis/special_functions/log/basic.lean b/src/analysis/special_functions/log/basic.lean index a4ffa13124f6f..6667ebe175a9b 100644 --- a/src/analysis/special_functions/log/basic.lean +++ b/src/analysis/special_functions/log/basic.lean @@ -325,3 +325,30 @@ lemma continuous_on.log (hf : continuous_on f s) (h₀ : ∀ x ∈ s, f x ≠ 0) λ x hx, (hf x hx).log (h₀ x hx) end continuity + + +section tendsto_comp_add_sub + +open filter +namespace real + +lemma tendsto_log_comp_add_sub_log (y : ℝ) : + tendsto (λ x:ℝ, log (x + y) - log x) at_top (𝓝 0) := +begin + refine tendsto.congr' (_ : ∀ᶠ (x : ℝ) in at_top, log (1 + y / x) = _) _, + { refine eventually.mp ((eventually_ne_at_top 0).and (eventually_gt_at_top (-y))) + (eventually_of_forall (λ x hx, _)), + rw ← log_div _ hx.1, + { congr' 1, + field_simp [hx.1] }, + { linarith [hx.2] } }, + { suffices : tendsto (λ (x : ℝ), log (1 + y / x)) at_top (𝓝 (log (1 + 0))), by simpa, + refine tendsto.log _ (by simp), + exact tendsto_const_nhds.add (tendsto_const_nhds.div_at_top tendsto_id) }, +end + +lemma tendsto_log_nat_add_one_sub_log : tendsto (λ (k : ℕ), log (k + 1) - log k) at_top (𝓝 0) := +(tendsto_log_comp_add_sub_log 1).comp tendsto_coe_nat_at_top_at_top + +end real +end tendsto_comp_add_sub diff --git a/src/data/real/conjugate_exponents.lean b/src/data/real/conjugate_exponents.lean index c11da7181c8e2..6661e83456482 100644 --- a/src/data/real/conjugate_exponents.lean +++ b/src/data/real/conjugate_exponents.lean @@ -112,4 +112,8 @@ lemma is_conjugate_exponent_conjugate_exponent {p : ℝ} (h : 1 < p) : p.is_conjugate_exponent (conjugate_exponent p) := (is_conjugate_exponent_iff h).2 rfl +lemma is_conjugate_exponent_one_div {a b : ℝ} (ha : 0 < a) (hb : 0 < b) (hab : a + b = 1) : + (1 / a).is_conjugate_exponent (1 / b) := +⟨by { rw [lt_div_iff ha, one_mul], linarith }, by { simp_rw one_div_one_div, exact hab }⟩ + end real From 68cc421841f2ebb8ad2b5a35a853895feb4b850a Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Fri, 20 Jan 2023 15:15:29 +0000 Subject: [PATCH 0306/1291] feat(data/finset/basic): `finset.to_list_eq_singleton_iff` (#18236) We add the finset analog of `multiset.to_list_eq_singleton_iff`. Mathlib4 pair: https://github.com/leanprover-community/mathlib4/pull/1726 --- src/data/finset/basic.lean | 8 ++++++++ 1 file changed, 8 insertions(+) diff --git a/src/data/finset/basic.lean b/src/data/finset/basic.lean index 8e804a6b8e8ae..28c672b401f2e 100644 --- a/src/data/finset/basic.lean +++ b/src/data/finset/basic.lean @@ -463,6 +463,9 @@ theorem not_mem_singleton {a b : α} : a ∉ ({b} : finset α) ↔ a ≠ b := no theorem mem_singleton_self (a : α) : a ∈ ({a} : finset α) := or.inl rfl +@[simp] lemma val_eq_singleton_iff {a : α} {s : finset α} : s.val = {a} ↔ s = {a} := +by { rw ←val_inj, refl } + lemma singleton_injective : injective (singleton : α → finset α) := λ a b h, mem_singleton.1 (h ▸ mem_singleton_self _) @@ -2215,6 +2218,11 @@ lemma coe_to_list (s : finset α) : (s.to_list : multiset α) = s.val := s.val.c @[simp] lemma to_list_to_finset [decidable_eq α] (s : finset α) : s.to_list.to_finset = s := by { ext, simp } +@[simp] lemma to_list_eq_singleton_iff {a : α} {s : finset α} : s.to_list = [a] ↔ s = {a} := +by rw [to_list, to_list_eq_singleton_iff, val_eq_singleton_iff] + +@[simp] lemma to_list_singleton : ∀ a, ({a} : finset α).to_list = [a] := to_list_singleton + lemma exists_list_nodup_eq [decidable_eq α] (s : finset α) : ∃ (l : list α), l.nodup ∧ l.to_finset = s := ⟨s.to_list, s.nodup_to_list, s.to_list_to_finset⟩ From 2445c98ae4b87eabebdde552593519b9b6dc350c Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 20 Jan 2023 18:38:42 +0000 Subject: [PATCH 0307/1291] fix(data/finsupp/basic): add missing `decidable` arguments in lemma statements (#18241) The resulting lemmas are more general than they were before. In order to ensure that this doesn't regress again, `open_locale classical` is now also removed from these files. Instead, we use the approach of: * Using the `classical` tactic in proofs * Using `by haveI := _; exact` in definitions, as `by classical; exact` leaks classicality up to the end of the next section. In a few places this means that `variables` lines have to be repeated on `def`s as Lean doesn't look inside tactic blocks to work out which variables are used. I also switched some anonymous constructors for named constructors in order to make the proof / data distinction a little easier to see. --- src/data/finsupp/basic.lean | 133 ++++++++++++++++++++++----------- src/data/finsupp/defs.lean | 128 ++++++++++++++++++------------- src/data/polynomial/basic.lean | 2 +- 3 files changed, 169 insertions(+), 94 deletions(-) diff --git a/src/data/finsupp/basic.lean b/src/data/finsupp/basic.lean index ae25e8f65ba2e..ad52601377153 100644 --- a/src/data/finsupp/basic.lean +++ b/src/data/finsupp/basic.lean @@ -45,7 +45,7 @@ This file is a `noncomputable theory` and uses classical logic throughout. noncomputable theory open finset function -open_locale classical big_operators +open_locale big_operators variables {α β γ ι M M' N P G H R S : Type*} @@ -84,12 +84,16 @@ lemma apply_eq_of_mem_graph {a : α} {m : M} {f : α →₀ M} (h : (a, m) ∈ f @[simp] lemma not_mem_graph_snd_zero (a : α) (f : α →₀ M) : (a, (0 : M)) ∉ f.graph := λ h, (mem_graph_iff.1 h).2.irrefl -@[simp] lemma image_fst_graph (f : α →₀ M) : f.graph.image prod.fst = f.support := -by simp only [graph, map_eq_image, image_image, embedding.coe_fn_mk, (∘), image_id'] +@[simp] lemma image_fst_graph [decidable_eq α] (f : α →₀ M) : f.graph.image prod.fst = f.support := +begin + classical, + simp only [graph, map_eq_image, image_image, embedding.coe_fn_mk, (∘), image_id'], +end lemma graph_injective (α M) [has_zero M] : injective (@graph α M _) := begin intros f g h, + classical, have hsup : f.support = g.support, by rw [← image_fst_graph, h, image_fst_graph], refine ext_iff'.2 ⟨hsup, λ x hx, apply_eq_of_mem_graph $ h.symm ▸ _⟩, exact mk_mem_graph _ (hsup ▸ hx) @@ -114,11 +118,15 @@ end { apply nodup_to_list } end⟩ -@[simp] lemma to_alist_keys_to_finset (f : α →₀ M) : f.to_alist.keys.to_finset = f.support := +@[simp] lemma to_alist_keys_to_finset [decidable_eq α] (f : α →₀ M) : + f.to_alist.keys.to_finset = f.support := by { ext, simp [to_alist, alist.mem_keys, alist.keys, list.keys] } @[simp] lemma mem_to_alist {f : α →₀ M} {x : α} : x ∈ f.to_alist ↔ f x ≠ 0 := -by rw [alist.mem_keys, ←list.mem_to_finset, to_alist_keys_to_finset, mem_support_iff] +begin + classical, + rw [alist.mem_keys, ←list.mem_to_finset, to_alist_keys_to_finset, mem_support_iff] +end end graph @@ -135,35 +143,44 @@ open list /-- Converts an association list into a finitely supported function via `alist.lookup`, sending absent keys to zero. -/ -@[simps] def lookup_finsupp (l : alist (λ x : α, M)) : α →₀ M := -{ support := (l.1.filter $ λ x, sigma.snd x ≠ 0).keys.to_finset, - to_fun := λ a, (l.lookup a).get_or_else 0, +def lookup_finsupp (l : alist (λ x : α, M)) : α →₀ M := +{ support := by haveI := classical.dec_eq α; haveI := classical.dec_eq M; exact + (l.1.filter $ λ x, sigma.snd x ≠ 0).keys.to_finset, + to_fun := λ a, by haveI := classical.dec_eq α; exact (l.lookup a).get_or_else 0, mem_support_to_fun := λ a, begin + classical, simp_rw [mem_to_finset, list.mem_keys, list.mem_filter, ←mem_lookup_iff], cases lookup a l; simp end } -alias lookup_finsupp_to_fun ← lookup_finsupp_apply +@[simp] lemma lookup_finsupp_apply [decidable_eq α] (l : alist (λ x : α, M)) (a : α) : + l.lookup_finsupp a = (l.lookup a).get_or_else 0 := +by convert rfl + +@[simp] lemma lookup_finsupp_support [decidable_eq α] [decidable_eq M] (l : alist (λ x : α, M)) : + l.lookup_finsupp.support = (l.1.filter $ λ x, sigma.snd x ≠ 0).keys.to_finset := +by convert rfl -lemma lookup_finsupp_eq_iff_of_ne_zero {l : alist (λ x : α, M)} {a : α} {x : M} (hx : x ≠ 0) : +lemma lookup_finsupp_eq_iff_of_ne_zero [decidable_eq α] + {l : alist (λ x : α, M)} {a : α} {x : M} (hx : x ≠ 0) : l.lookup_finsupp a = x ↔ x ∈ l.lookup a := -by { rw lookup_finsupp_to_fun, cases lookup a l with m; simp [hx.symm] } +by { rw lookup_finsupp_apply, cases lookup a l with m; simp [hx.symm] } -lemma lookup_finsupp_eq_zero_iff {l : alist (λ x : α, M)} {a : α} : +lemma lookup_finsupp_eq_zero_iff [decidable_eq α] {l : alist (λ x : α, M)} {a : α} : l.lookup_finsupp a = 0 ↔ a ∉ l ∨ (0 : M) ∈ l.lookup a := -by { rw [lookup_finsupp_to_fun, ←lookup_eq_none], cases lookup a l with m; simp } +by { rw [lookup_finsupp_apply, ←lookup_eq_none], cases lookup a l with m; simp } @[simp] lemma empty_lookup_finsupp : lookup_finsupp (∅ : alist (λ x : α, M)) = 0 := -by { ext, simp } +by { classical, ext, simp } -@[simp] lemma insert_lookup_finsupp (l : alist (λ x : α, M)) (a : α) (m : M) : +@[simp] lemma insert_lookup_finsupp [decidable_eq α] (l : alist (λ x : α, M)) (a : α) (m : M) : (l.insert a m).lookup_finsupp = l.lookup_finsupp.update a m := by { ext b, by_cases h : b = a; simp [h] } @[simp] lemma singleton_lookup_finsupp (a : α) (m : M) : (singleton a m).lookup_finsupp = finsupp.single a m := -by simp [←alist.insert_empty] +by { classical, simp [←alist.insert_empty] } @[simp] lemma _root_.finsupp.to_alist_lookup_finsupp (f : α →₀ M) : f.to_alist.lookup_finsupp = f := begin @@ -370,7 +387,11 @@ lemma equiv_map_domain_trans' (f : α ≃ β) (g : β ≃ γ) : @[simp] lemma equiv_map_domain_single (f : α ≃ β) (a : α) (b : M) : equiv_map_domain f (single a b) = single (f a) b := -by ext x; simp only [single_apply, equiv.apply_eq_iff_eq_symm_apply, equiv_map_domain_apply]; congr +begin + classical, + ext x, + simp only [single_apply, equiv.apply_eq_iff_eq_symm_apply, equiv_map_domain_apply], +end @[simp] lemma equiv_map_domain_zero {f : α ≃ β} : equiv_map_domain f (0 : α →₀ M) = (0 : β →₀ M) := by ext x; simp only [equiv_map_domain_apply, coe_zero, pi.zero_apply] @@ -537,6 +558,7 @@ lemma map_domain_apply' (S : set α) {f : α → β} (x : α →₀ M) (hS : (x.support : set α) ⊆ S) (hf : set.inj_on f S) {a : α} (ha : a ∈ S) : map_domain f x (f a) = x a := begin + classical, rw [map_domain, sum_apply, sum], simp_rw single_apply, by_cases hax : a ∈ x.support, @@ -646,6 +668,7 @@ lemma map_domain_inj_on (S : set α) {f : α → β} begin intros v₁ hv₁ v₂ hv₂ eq, ext a, + classical, by_cases h : a ∈ v₁.support ∪ v₂.support, { rw [← map_domain_apply' S _ hv₁ hf _, ← map_domain_apply' S _ hv₂ hf _, eq]; { apply set.union_subset hv₁ hv₂, @@ -796,7 +819,7 @@ by { ext, simp, } @[simp] lemma some_single_some [has_zero M] (a : α) (m : M) : (single (option.some a) m : option α →₀ M).some = single a m := -by { ext b, simp [single_apply], } +by { classical, ext b, simp [single_apply], } @[to_additive] lemma prod_option_index [add_comm_monoid M] [comm_monoid N] @@ -831,8 +854,8 @@ variables [has_zero M] (p : α → Prop) (f : α →₀ M) /-- `filter p f` is the finitely supported function that is `f a` if `p a` is true and 0 otherwise. -/ def filter (p : α → Prop) (f : α →₀ M) : α →₀ M := -{ to_fun := λ a, if p a then f a else 0, - support := f.support.filter (λ a, p a), +{ to_fun := λ a, by haveI := classical.dec_pred p; exact if p a then f a else 0, + support := by haveI := classical.dec_pred p; exact f.support.filter (λ a, p a), mem_support_to_fun := λ a, by split_ifs; { simp only [h, mem_filter, mem_support_iff], tauto } } lemma filter_apply (a : α) [D : decidable (p a)] : f.filter p a = if p a then f a else 0 := @@ -849,16 +872,16 @@ by simp only [fun_like.ext_iff, filter_eq_indicator, set.indicator_apply_eq_self not_imp_comm] @[simp] lemma filter_apply_pos {a : α} (h : p a) : f.filter p a = f a := -if_pos h +by { classical, convert if_pos h } @[simp] lemma filter_apply_neg {a : α} (h : ¬ p a) : f.filter p a = 0 := -if_neg h +by { classical, convert if_neg h } @[simp] lemma support_filter [D : decidable_pred p] : (f.filter p).support = f.support.filter p := by rw subsingleton.elim D; refl lemma filter_zero : (0 : α →₀ M).filter p = 0 := -by rw [← support_eq_empty, support_filter, support_zero, finset.filter_empty] +by { classical, rw [← support_eq_empty, support_filter, support_zero, finset.filter_empty] } @[simp] lemma filter_single_of_pos {a : α} {b : M} (h : p a) : (single a b).filter p = single a b := @@ -870,6 +893,7 @@ by rw [← support_eq_empty, support_filter, support_zero, finset.filter_empty] @[to_additive] lemma prod_filter_index [comm_monoid N] (g : α → M → N) : (f.filter p).prod g = ∏ x in (f.filter p).support, g x (f x) := begin + classical, refine finset.prod_congr rfl (λ x hx, _), rw [support_filter, finset.mem_filter] at hx, rw [filter_apply_pos _ _ hx.2] @@ -877,7 +901,10 @@ end @[simp, to_additive] lemma prod_filter_mul_prod_filter_not [comm_monoid N] (g : α → M → N) : (f.filter p).prod g * (f.filter (λ a, ¬ p a)).prod g = f.prod g := -by simp_rw [prod_filter_index, support_filter, prod_filter_mul_prod_filter_not, finsupp.prod] +begin + classical, + simp_rw [prod_filter_index, support_filter, prod_filter_mul_prod_filter_not, finsupp.prod] +end @[simp, to_additive] lemma prod_div_prod_filter [comm_group G] (g : α → M → G) : f.prod g / (f.filter p).prod g = (f.filter (λ a, ¬p a)).prod g := @@ -897,11 +924,12 @@ section frange variables [has_zero M] /-- `frange f` is the image of `f` on the support of `f`. -/ -def frange (f : α →₀ M) : finset M := finset.image f f.support +def frange (f : α →₀ M) : finset M := +by haveI := classical.dec_eq M; exact finset.image f f.support theorem mem_frange {f : α →₀ M} {y : M} : y ∈ f.frange ↔ y ≠ 0 ∧ ∃ x, f x = y := -finset.mem_image.trans +by classical; exact finset.mem_image.trans ⟨λ ⟨x, hx1, hx2⟩, ⟨hx2 ▸ mem_support_iff.1 hx1, x, hx2⟩, λ ⟨hy, x, hx⟩, ⟨x, mem_support_iff.2 (hx.symm ▸ hy), hx⟩⟩ @@ -909,8 +937,13 @@ theorem zero_not_mem_frange {f : α →₀ M} : (0:M) ∉ f.frange := λ H, (mem_frange.1 H).1 rfl theorem frange_single {x : α} {y : M} : frange (single x y) ⊆ {y} := -λ r hr, let ⟨t, ht1, ht2⟩ := mem_frange.1 hr in ht2 ▸ - (by rw single_apply at ht2 ⊢; split_ifs at ht2 ⊢; [exact finset.mem_singleton_self _, cc]) +λ r hr, let ⟨t, ht1, ht2⟩ := mem_frange.1 hr in ht2 ▸ begin + classical, + rw single_apply at ht2 ⊢, + split_ifs at ht2 ⊢, + { exact finset.mem_singleton_self _ }, + { exact (t ht2.symm).elim } +end end frange @@ -925,7 +958,9 @@ variables [has_zero M] {p : α → Prop} /-- `subtype_domain p f` is the restriction of the finitely supported function `f` to subtype `p`. -/ def subtype_domain (p : α → Prop) (f : α →₀ M) : (subtype p →₀ M) := -⟨f.support.subtype p, f ∘ coe, λ a, by simp only [mem_subtype, mem_support_iff]⟩ +{ support := by haveI := classical.dec_pred p; exact f.support.subtype p, + to_fun := f ∘ coe, + mem_support_to_fun := λ a, by simp only [mem_subtype, mem_support_iff] } @[simp] lemma support_subtype_domain [D : decidable_pred p] {f : α →₀ M} : (subtype_domain p f).support = f.support.subtype p := @@ -940,19 +975,23 @@ rfl lemma subtype_domain_eq_zero_iff' {f : α →₀ M} : f.subtype_domain p = 0 ↔ ∀ x, p x → f x = 0 := -by simp_rw [← support_eq_empty, support_subtype_domain, subtype_eq_empty, not_mem_support_iff] +begin + classical, + simp_rw [← support_eq_empty, support_subtype_domain, subtype_eq_empty, not_mem_support_iff] +end lemma subtype_domain_eq_zero_iff {f : α →₀ M} (hf : ∀ x ∈ f.support , p x) : f.subtype_domain p = 0 ↔ f = 0 := subtype_domain_eq_zero_iff'.trans ⟨λ H, ext $ λ x, - if hx : p x then H x hx else not_mem_support_iff.1 $ mt (hf x) hx, λ H x _, by simp [H]⟩ + by classical; exact + if hx : p x then H x hx else not_mem_support_iff.1 $ mt (hf x) hx, λ H x _, by simp [H]⟩ @[to_additive] lemma prod_subtype_domain_index [comm_monoid N] {v : α →₀ M} {h : α → M → N} (hp : ∀x∈v.support, p x) : (v.subtype_domain p).prod (λa b, h a b) = v.prod h := prod_bij (λp _, p.val) - (λ _, mem_subtype.1) + (λ _, by classical; exact mem_subtype.1) (λ _ _, rfl) (λ _ _ _ _, subtype.eq) (λ b hb, ⟨⟨b, hp b hb⟩, mem_subtype.2 hb, rfl⟩) @@ -1075,6 +1114,7 @@ f.sum $ λp c, single p.1 (single p.2 c) @[simp] lemma curry_apply (f : (α × β) →₀ M) (x : α) (y : β) : f.curry x y = f (x, y) := begin + classical, have : ∀ (b : α × β), single b.fst (single b.snd (f b)) x y = if b = (x, y) then f b else 0, { rintros ⟨b₁, b₂⟩, simp [single_apply, ite_apply, prod.ext_iff, ite_and], @@ -1115,6 +1155,7 @@ by refine ⟨finsupp.curry, finsupp.uncurry, λ f, _, λ f, _⟩; simp only [ lemma filter_curry (f : α × β →₀ M) (p : α → Prop) : (f.filter (λa:α×β, p a.1)).curry = f.curry.filter p := begin + classical, rw [finsupp.curry, finsupp.curry, finsupp.sum, finsupp.sum, filter_sum, support_filter, sum_filter], refine finset.sum_congr rfl _, @@ -1143,7 +1184,8 @@ section sum def sum_elim {α β γ : Type*} [has_zero γ] (f : α →₀ γ) (g : β →₀ γ) : α ⊕ β →₀ γ := on_finset - ((f.support.map ⟨_, sum.inl_injective⟩) ∪ g.support.map ⟨_, sum.inr_injective⟩) + (by haveI := classical.dec_eq α; haveI := classical.dec_eq β; + exact (f.support.map ⟨_, sum.inl_injective⟩) ∪ g.support.map ⟨_, sum.inr_injective⟩) (sum.elim f g) (λ ab h, by { cases ab with a b; simp only [sum.elim_inl, sum.elim_inr] at h; simpa }) @@ -1475,12 +1517,15 @@ between the subtype of finitely supported functions with support contained in `s the type of finitely supported functions from `s`. -/ def restrict_support_equiv (s : set α) (M : Type*) [add_comm_monoid M] : {f : α →₀ M // ↑f.support ⊆ s } ≃ (s →₀ M) := -begin - refine ⟨λf, subtype_domain (λx, x ∈ s) f.1, λ f, ⟨f.map_domain subtype.val, _⟩, _, _⟩, - { refine set.subset.trans (finset.coe_subset.2 map_domain_support) _, +{ to_fun := λ f, subtype_domain (λ x, x ∈ s) f.1, + inv_fun := λ f, ⟨f.map_domain subtype.val, begin + classical, + refine set.subset.trans (finset.coe_subset.2 map_domain_support) _, rw [finset.coe_image, set.image_subset_iff], - exact assume x hx, x.2 }, - { rintros ⟨f, hf⟩, + exact assume x hx, x.2, + end⟩, + left_inv := begin + rintros ⟨f, hf⟩, apply subtype.eq, ext a, dsimp only, @@ -1490,12 +1535,13 @@ begin { convert map_domain_notin_range _ _ h, rw [← not_mem_support_iff], refine mt _ h, - exact assume ha, ⟨⟨a, hf ha⟩, rfl⟩ } }, - { assume f, + exact assume ha, ⟨⟨a, hf ha⟩, rfl⟩ } + end, + right_inv := λ f, begin ext ⟨a, ha⟩, dsimp only, - rw [subtype_domain_apply, map_domain_apply subtype.val_injective] } -end + rw [subtype_domain_apply, map_domain_apply subtype.val_injective] + end } /-- Given `add_comm_monoid M` and `e : α ≃ β`, `dom_congr e` is the corresponding `equiv` between `α →₀ M` and `β →₀ M`. @@ -1556,7 +1602,8 @@ end /-- Given `l`, a finitely supported function from the sigma type `Σ (i : ι), αs i` to `β`, `split_support l` is the finset of indices in `ι` that appear in the support of `l`. -/ -def split_support : finset ι := l.support.image sigma.fst +def split_support (l : (Σ i, αs i) →₀ M) : finset ι := +by haveI := classical.dec_eq ι; exact l.support.image sigma.fst lemma mem_split_support_iff_nonzero (i : ι) : i ∈ split_support l ↔ split l i ≠ 0 := diff --git a/src/data/finsupp/defs.lean b/src/data/finsupp/defs.lean index ac9ae1845a73b..da86462d707b7 100644 --- a/src/data/finsupp/defs.lean +++ b/src/data/finsupp/defs.lean @@ -80,7 +80,7 @@ This file is a `noncomputable theory` and uses classical logic throughout. noncomputable theory open finset function -open_locale classical big_operators +open_locale big_operators variables {α β γ ι M M' N P G H R S : Type*} @@ -146,7 +146,7 @@ by rw [← coe_zero, coe_fn_inj] lemma ext_iff' {f g : α →₀ M} : f = g ↔ f.support = g.support ∧ ∀ x ∈ f.support, f x = g x := ⟨λ h, h ▸ ⟨rfl, λ _ _, rfl⟩, λ ⟨h₁, h₂⟩, ext $ λ a, - if h : a ∈ f.support then h₂ a h else + by classical; exact if h : a ∈ f.support then h₂ a h else have hf : f a = 0, from not_mem_support_iff.1 h, have hg : g a = 0, by rwa [h₁, not_mem_support_iff] at h, by rw [hf, hg]⟩ @@ -202,14 +202,17 @@ variables [has_zero M] {a a' : α} {b : M} /-- `single a b` is the finitely supported function with value `b` at `a` and zero otherwise. -/ def single (a : α) (b : M) : α →₀ M := -⟨if b = 0 then ∅ else {a}, pi.single a b, λ a', begin - obtain rfl | hb := eq_or_ne b 0, - { simp }, - rw [if_neg hb, mem_singleton], - obtain rfl | ha := eq_or_ne a' a, - { simp [hb] }, - simp [pi.single_eq_of_ne', ha], -end⟩ +{ support := by haveI := classical.dec_eq M; exact if b = 0 then ∅ else {a}, + to_fun := by haveI := classical.dec_eq α; exact pi.single a b, + mem_support_to_fun := λ a', begin + classical, + obtain rfl | hb := eq_or_ne b 0, + { simp }, + rw [if_neg hb, mem_singleton], + obtain rfl | ha := eq_or_ne a' a, + { simp [hb] }, + simp [pi.single_eq_of_ne', ha], + end } lemma single_apply [decidable (a = a')] : single a b a' = if a = a' then b else 0 := by { simp_rw [@eq_comm _ a a'], convert pi.single_apply _ _ _, } @@ -401,26 +404,29 @@ If `b = 0`, this amounts to removing `a` from the `finsupp.support`. Otherwise, if `a` was not in the `finsupp.support`, it is added to it. This is the finitely-supported version of `function.update`. -/ -def update : α →₀ M := -⟨if b = 0 then f.support.erase a else insert a f.support, - function.update f a b, - λ i, begin +def update (f : α →₀ M) (a : α) (b : M) : α →₀ M := +{ support := by haveI := classical.dec_eq α; haveI := classical.dec_eq M; exact + if b = 0 then f.support.erase a else insert a f.support, + to_fun := by haveI := classical.dec_eq α; exact + function.update f a b, + mem_support_to_fun := λ i, begin simp only [function.update_apply, ne.def], split_ifs with hb ha ha hb; simp [ha, hb] - end⟩ + end } @[simp] lemma coe_update [decidable_eq α] : (f.update a b : α → M) = function.update f a b := by convert rfl @[simp] lemma update_self : f.update a (f a) = f := -by { ext, simp } +by { classical, ext, simp } @[simp] lemma zero_update : update 0 a b = single a b := -by { ext, rw single_eq_update, refl } +by { classical, ext, rw single_eq_update, refl } -lemma support_update [decidable_eq α] : support (f.update a b) = - if b = 0 then f.support.erase a else insert a f.support := by convert rfl +lemma support_update [decidable_eq α] [decidable_eq M] : + support (f.update a b) = if b = 0 then f.support.erase a else insert a f.support := +by convert rfl @[simp] lemma support_update_zero [decidable_eq α] : support (f.update a 0) = f.support.erase a := by convert if_pos rfl @@ -428,7 +434,7 @@ lemma support_update [decidable_eq α] : support (f.update a b) = variables {b} lemma support_update_ne_zero [decidable_eq α] (h : b ≠ 0) : - support (f.update a b) = insert a f.support := by convert if_neg h + support (f.update a b) = insert a f.support := by { classical, convert if_neg h } end update @@ -443,20 +449,21 @@ variables [has_zero M] If `a` is not in the support of `f` then `erase a f = f`. -/ def erase (a : α) (f : α →₀ M) : α →₀ M := -⟨f.support.erase a, (λa', if a' = a then 0 else f a'), - assume a', by rw [mem_erase, mem_support_iff]; split_ifs; +{ support := by haveI := classical.dec_eq α; exact f.support.erase a, + to_fun := λ a', by haveI := classical.dec_eq α; exact if a' = a then 0 else f a', + mem_support_to_fun := assume a', by rw [mem_erase, mem_support_iff]; split_ifs; [exact ⟨λ H _, H.1 h, λ H, (H rfl).elim⟩, - exact and_iff_right h]⟩ + exact and_iff_right h] } @[simp] lemma support_erase [decidable_eq α] {a : α} {f : α →₀ M} : (f.erase a).support = f.support.erase a := by convert rfl @[simp] lemma erase_same {a : α} {f : α →₀ M} : (f.erase a) a = 0 := -if_pos rfl +by convert if_pos rfl @[simp] lemma erase_ne {a a' : α} {f : α →₀ M} (h : a' ≠ a) : (f.erase a) a' = f a' := -if_neg h +by { classical, convert if_neg h } @[simp] lemma erase_single {a : α} {b : M} : (erase a (single a b)) = 0 := begin @@ -480,7 +487,7 @@ begin end @[simp] lemma erase_zero (a : α) : erase a (0 : α →₀ M) = 0 := -by rw [← support_eq_empty, support_erase, support_zero, erase_empty] +by { classical, rw [← support_eq_empty, support_erase, support_zero, erase_empty] } end erase @@ -493,7 +500,9 @@ variables [has_zero M] The function must be `0` outside of `s`. Use this when the set needs to be filtered anyways, otherwise a better set representation is often available. -/ def on_finset (s : finset α) (f : α → M) (hf : ∀a, f a ≠ 0 → a ∈ s) : α →₀ M := -⟨s.filter (λa, f a ≠ 0), f, by simpa⟩ +{ support := by haveI := classical.dec_eq M; exact s.filter (λa, f a ≠ 0), + to_fun := f, + mem_support_to_fun := by simpa } @[simp] lemma on_finset_apply {s : finset α} {f : α → M} {hf a} : (on_finset s f hf : α →₀ M) a = f a := @@ -501,17 +510,17 @@ rfl @[simp] lemma support_on_finset_subset {s : finset α} {f : α → M} {hf} : (on_finset s f hf).support ⊆ s := -filter_subset _ _ +by convert filter_subset _ _ @[simp] lemma mem_support_on_finset {s : finset α} {f : α → M} (hf : ∀ (a : α), f a ≠ 0 → a ∈ s) {a : α} : a ∈ (finsupp.on_finset s f hf).support ↔ f a ≠ 0 := by rw [finsupp.mem_support_iff, finsupp.on_finset_apply] -lemma support_on_finset +lemma support_on_finset [decidable_eq M] {s : finset α} {f : α → M} (hf : ∀ (a : α), f a ≠ 0 → a ∈ s) : (finsupp.on_finset s f hf).support = s.filter (λ a, f a ≠ 0) := -rfl +by convert rfl end on_finset @@ -577,7 +586,10 @@ support_on_finset_subset @[simp] lemma map_range_single {f : M → N} {hf : f 0 = 0} {a : α} {b : M} : map_range f hf (single a b) = single a (f b) := -ext $ λ a', by simpa only [single_eq_pi_single] using pi.apply_single _ (λ _, hf) a _ a' +ext $ λ a', begin + classical, + simpa only [single_eq_pi_single] using pi.apply_single _ (λ _, hf) a _ a' +end lemma support_map_range_of_injective {e : M → N} (he0 : e 0 = 0) (f : ι →₀ M) (he : function.injective e) : @@ -599,18 +611,22 @@ variables [has_zero M] [has_zero N] is the finitely supported function whose value at `f a : β` is `v a`. For a `b : β` outside the range of `f`, it is zero. -/ def emb_domain (f : α ↪ β) (v : α →₀ M) : β →₀ M := -begin - refine ⟨v.support.map f, λa₂, - if h : a₂ ∈ v.support.map f then v (v.support.choose (λa₁, f a₁ = a₂) _) else 0, _⟩, - { rcases finset.mem_map.1 h with ⟨a, ha, rfl⟩, - exact exists_unique.intro a ⟨ha, rfl⟩ (assume b ⟨_, hb⟩, f.injective hb) }, - { assume a₂, +{ support := v.support.map f, + to_fun := λ a₂, + by haveI := classical.dec_eq β; exact + if h : a₂ ∈ v.support.map f + then v (v.support.choose (λa₁, f a₁ = a₂) begin + rcases finset.mem_map.1 h with ⟨a, ha, rfl⟩, + exact exists_unique.intro a ⟨ha, rfl⟩ (assume b ⟨_, hb⟩, f.injective hb) + end) + else 0, + mem_support_to_fun := λ a₂, begin split_ifs, { simp only [h, true_iff, ne.def], rw [← not_mem_support_iff, not_not], apply finset.choose_mem }, - { simp only [h, ne.def, ne_self_iff_false] } } -end + { simp only [h, ne.def, ne_self_iff_false] } + end } @[simp] lemma support_emb_domain (f : α ↪ β) (v : α →₀ M) : (emb_domain f v).support = v.support.map f := @@ -622,6 +638,7 @@ rfl @[simp] lemma emb_domain_apply (f : α ↪ β) (v : α →₀ M) (a : α) : emb_domain f v (f a) = v a := begin + classical, change dite _ _ _ = _, split_ifs; rw [finset.mem_map' f] at h, { refine congr_arg (v : α → M) (f.inj' _), @@ -632,6 +649,7 @@ end lemma emb_domain_notin_range (f : α ↪ β) (v : α →₀ M) (a : β) (h : a ∉ set.range f) : emb_domain f v a = 0 := begin + classical, refine dif_neg (mt (assume h, _) h), rcases finset.mem_map.1 h with ⟨a, h, rfl⟩, exact set.mem_range_self a @@ -665,6 +683,7 @@ lemma single_of_emb_domain_single (h : l.emb_domain f = single a b) : ∃ x, l = single x b ∧ f x = a := begin + classical, have h_map_support : finset.map f (l.support) = {a}, by rw [←support_emb_domain, h, support_single_ne_zero _ hb]; refl, have ha : a ∈ finset.map f (l.support), @@ -685,6 +704,7 @@ end @[simp] lemma emb_domain_single (f : α ↪ β) (a : α) (m : M) : emb_domain f (single a m) = single (f a) m := begin + classical, ext b, by_cases h : b ∈ set.range f, { rcases h with ⟨a', rfl⟩, @@ -706,11 +726,13 @@ variables [has_zero M] [has_zero N] [has_zero P] `zip_with f hf g₁ g₂` is the finitely supported function `α →₀ P` satisfying `zip_with f hf g₁ g₂ a = f (g₁ a) (g₂ a)`, which is well-defined when `f 0 0 = 0`. -/ def zip_with (f : M → N → P) (hf : f 0 0 = 0) (g₁ : α →₀ M) (g₂ : α →₀ N) : α →₀ P := -on_finset (g₁.support ∪ g₂.support) (λa, f (g₁ a) (g₂ a)) $ λ a H, -begin - simp only [mem_union, mem_support_iff, ne], rw [← not_and_distrib], - rintro ⟨h₁, h₂⟩, rw [h₁, h₂] at H, exact H hf -end +on_finset + (by haveI := classical.dec_eq α; exact g₁.support ∪ g₂.support) + (λa, f (g₁ a) (g₂ a)) $ λ a H, + begin + simp only [mem_union, mem_support_iff, ne], rw [← not_and_distrib], + rintro ⟨h₁, h₂⟩, rw [h₁, h₂] at H, exact H hf + end @[simp] lemma zip_with_apply {f : M → N → P} {hf : f 0 0 = 0} {g₁ : α →₀ M} {g₂ : α →₀ N} {a : α} : @@ -782,6 +804,7 @@ noncomputable def coe_fn_add_hom : (α →₀ M) →+ (α → M) := lemma update_eq_single_add_erase (f : α →₀ M) (a : α) (b : M) : f.update a b = single a b + f.erase a := begin + classical, ext j, rcases eq_or_ne a j with rfl|h, { simp }, @@ -791,6 +814,7 @@ end lemma update_eq_erase_add_single (f : α →₀ M) (a : α) (b : M) : f.update a b = f.erase a + single a b := begin + classical, ext j, rcases eq_or_ne a j with rfl|h, { simp }, @@ -820,30 +844,33 @@ protected theorem induction {p : (α →₀ M) → Prop} (f : α →₀ M) (h0 : p 0) (ha : ∀a b (f : α →₀ M), a ∉ f.support → b ≠ 0 → p f → p (single a b + f)) : p f := suffices ∀s (f : α →₀ M), f.support = s → p f, from this _ _ rfl, -assume s, finset.induction_on s (λ f hf, by rwa [support_eq_empty.1 hf]) $ +assume s, finset.cons_induction_on s (λ f hf, by rwa [support_eq_empty.1 hf]) $ assume a s has ih f hf, suffices p (single a (f a) + f.erase a), by rwa [single_add_erase] at this, begin + classical, apply ha, { rw [support_erase, mem_erase], exact λ H, H.1 rfl }, - { rw [← mem_support_iff, hf], exact mem_insert_self _ _ }, + { rw [← mem_support_iff, hf], exact mem_cons_self _ _ }, { apply ih _ _, - rw [support_erase, hf, finset.erase_insert has] } + rw [support_erase, hf, finset.erase_cons] } end lemma induction₂ {p : (α →₀ M) → Prop} (f : α →₀ M) (h0 : p 0) (ha : ∀a b (f : α →₀ M), a ∉ f.support → b ≠ 0 → p f → p (f + single a b)) : p f := suffices ∀s (f : α →₀ M), f.support = s → p f, from this _ _ rfl, -assume s, finset.induction_on s (λ f hf, by rwa [support_eq_empty.1 hf]) $ +assume s, finset.cons_induction_on s (λ f hf, by rwa [support_eq_empty.1 hf]) $ assume a s has ih f hf, suffices p (f.erase a + single a (f a)), by rwa [erase_add_single] at this, begin + classical, apply ha, { rw [support_erase, mem_erase], exact λ H, H.1 rfl }, - { rw [← mem_support_iff, hf], exact mem_insert_self _ _ }, + { rw [← mem_support_iff, hf], + exact mem_cons_self _ _ }, { apply ih _ _, - rw [support_erase, hf, finset.erase_insert has] } + rw [support_erase, hf, finset.erase_cons] } end lemma induction_linear {p : (α →₀ M) → Prop} (f : α →₀ M) @@ -988,6 +1015,7 @@ lemma single_add_single_eq_single_add_single [add_comm_monoid M] single k u + single l v = single m u + single n v ↔ (k = m ∧ l = n) ∨ (u = v ∧ k = n ∧ l = m) ∨ (u + v = 0 ∧ k = l ∧ m = n) := begin + classical, simp_rw [fun_like.ext_iff, coe_add, single_eq_pi_single, ←funext_iff], exact pi.single_add_single_eq_single_add_single hu hv, end diff --git a/src/data/polynomial/basic.lean b/src/data/polynomial/basic.lean index 8b825846cbcb2..1bd981b8111b8 100644 --- a/src/data/polynomial/basic.lean +++ b/src/data/polynomial/basic.lean @@ -781,7 +781,7 @@ by { ext, rw [coeff_update_apply, coeff_erase] } lemma support_update (p : R[X]) (n : ℕ) (a : R) [decidable (a = 0)] : support (p.update n a) = if a = 0 then p.support.erase n else insert n p.support := -by { cases p, simp only [support, update, support_update], congr } +by { classical, cases p, simp only [support, update, support_update], congr } lemma support_update_zero (p : R[X]) (n : ℕ) : support (p.update n 0) = p.support.erase n := From 2d3f0c84d58173336353bb4a1fe8446f0fdb315f Mon Sep 17 00:00:00 2001 From: themathqueen <2497250a@research.gla.ac.uk> Date: Fri, 20 Jan 2023 21:58:55 +0000 Subject: [PATCH 0308/1291] feat(linear_algebra/dual): lemmas (#18228) some lemmas --- src/linear_algebra/dual.lean | 45 ++++++++++++++++++++++++++++++++++++ 1 file changed, 45 insertions(+) diff --git a/src/linear_algebra/dual.lean b/src/linear_algebra/dual.lean index 425bd69a89e60..632e5effca548 100644 --- a/src/linear_algebra/dual.lean +++ b/src/linear_algebra/dual.lean @@ -854,6 +854,16 @@ begin rw [dual_map_apply, hx, map_zero] end +/-- if a linear map is surjective, then its dual is injective -/ +lemma dual_map_injective_of_surjective {f : M₁ →ₗ[R] M₂} (hf : function.surjective f) : + function.injective f.dual_map := +begin + obtain ⟨g, hg⟩ := hf.has_right_inverse, + intros x y hxy, + ext w, + simpa only [dual_map_apply, hg w] using fun_like.congr_fun hxy (g w), +end + section finite_dimensional variables {K : Type*} [field K] {V₁ : Type*} {V₂ : Type*} @@ -886,6 +896,41 @@ begin exact finrank_range_add_finrank_ker f, end +/-- `f.dual_map` is injective if and only if `f` is surjective -/ +@[simp] lemma dual_map_injective_iff {f : V₁ →ₗ[K] V₂} : + function.injective f.dual_map ↔ function.surjective f := +begin + refine ⟨_, λ h, dual_map_injective_of_surjective h⟩, + rw [← range_eq_top, ← ker_eq_bot], + intro h, + apply finite_dimensional.eq_top_of_finrank_eq, + rw ← finrank_eq_zero at h, + rw [← add_zero (finite_dimensional.finrank K f.range), ← h, + ← linear_map.finrank_range_dual_map_eq_finrank_range, + linear_map.finrank_range_add_finrank_ker, subspace.dual_finrank_eq], +end + +/-- `f.dual_map` is surjective if and only if `f` is injective -/ +@[simp] lemma dual_map_surjective_iff [finite_dimensional K V₁] {f : V₁ →ₗ[K] V₂} : + function.surjective f.dual_map ↔ function.injective f := +begin + rw [← range_eq_top, ← ker_eq_bot], + split, + { intro h, + rw [← finrank_eq_zero, ← add_right_inj (finite_dimensional.finrank K f.range), + add_zero, linear_map.finrank_range_add_finrank_ker, + ← linear_map.finrank_range_dual_map_eq_finrank_range, h, + finrank_top, subspace.dual_finrank_eq], }, + { intro h, + rw [range_dual_map_eq_dual_annihilator_ker, h], + exact submodule.dual_annihilator_bot, }, +end + +/-- `f.dual_map` is bijective if and only if `f` is -/ +@[simp] lemma dual_map_bijective_iff [finite_dimensional K V₁] {f : V₁ →ₗ[K] V₂} : + function.bijective f.dual_map ↔ function.bijective f := +by simp_rw [function.bijective, dual_map_surjective_iff, dual_map_injective_iff, and.comm] + end finite_dimensional section field From a63928c34ec358b5edcda2bf7513c50052a5230f Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Fri, 20 Jan 2023 23:55:54 +0000 Subject: [PATCH 0309/1291] chore(analysis/convex/topology): split (#18187) Split the facts about the topology of convex sets into two files, for topological vector spaces and normed spaces. This makes the former a more elementary file, as indeed it feels like it should be. It also slightly decreases the length of the longest path in mathlib. --- src/analysis/calculus/mean_value.lean | 2 +- .../complex/upper_half_plane/topology.lean | 2 +- src/analysis/convex/normed.lean | 143 +++++++++++++++++ src/analysis/convex/side.lean | 2 +- src/analysis/convex/strict_convex_space.lean | 2 +- src/analysis/convex/topology.lean | 148 ++---------------- src/analysis/special_functions/bernstein.lean | 1 + src/topology/algebra/module/basic.lean | 21 +++ .../algebra/module/finite_dimension.lean | 23 --- src/topology/metric_space/metrizable.lean | 1 + src/topology/tietze_extension.lean | 1 + 11 files changed, 181 insertions(+), 165 deletions(-) create mode 100644 src/analysis/convex/normed.lean diff --git a/src/analysis/calculus/mean_value.lean b/src/analysis/calculus/mean_value.lean index f64fd61e91c71..636a38d50cb4e 100644 --- a/src/analysis/calculus/mean_value.lean +++ b/src/analysis/calculus/mean_value.lean @@ -5,7 +5,7 @@ Authors: Sébastien Gouëzel, Yury Kudryashov -/ import analysis.calculus.local_extr import analysis.convex.slope -import analysis.convex.topology +import analysis.convex.normed import data.complex.is_R_or_C /-! diff --git a/src/analysis/complex/upper_half_plane/topology.lean b/src/analysis/complex/upper_half_plane/topology.lean index a47570b6e516b..65fba4c6df834 100644 --- a/src/analysis/complex/upper_half_plane/topology.lean +++ b/src/analysis/complex/upper_half_plane/topology.lean @@ -5,7 +5,7 @@ Authors: Yury G. Kudryashov -/ import analysis.complex.upper_half_plane.basic import analysis.convex.contractible -import analysis.convex.topology +import analysis.convex.normed import analysis.convex.complex import analysis.complex.re_im_topology import topology.homotopy.contractible diff --git a/src/analysis/convex/normed.lean b/src/analysis/convex/normed.lean new file mode 100644 index 0000000000000..b5ac0d553abd5 --- /dev/null +++ b/src/analysis/convex/normed.lean @@ -0,0 +1,143 @@ +/- +Copyright (c) 2020 Yury Kudryashov. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Alexander Bentkamp, Yury Kudryashov +-/ +import analysis.convex.jensen +import analysis.convex.topology +import analysis.normed.group.pointwise +import analysis.normed_space.ray + +/-! +# Topological and metric properties of convex sets in normed spaces + +We prove the following facts: + +* `convex_on_norm`, `convex_on_dist` : norm and distance to a fixed point is convex on any convex + set; +* `convex_on_univ_norm`, `convex_on_univ_dist` : norm and distance to a fixed point is convex on + the whole space; +* `convex_hull_ediam`, `convex_hull_diam` : convex hull of a set has the same (e)metric diameter + as the original set; +* `bounded_convex_hull` : convex hull of a set is bounded if and only if the original set + is bounded. +* `bounded_std_simplex`, `is_closed_std_simplex`, `compact_std_simplex`: topological properties + of the standard simplex. +-/ + +variables {ι : Type*} {E : Type*} + +open metric set +open_locale pointwise convex + +variables [seminormed_add_comm_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 (hs : convex ℝ s) : convex_on ℝ s norm := +⟨hs, λ x hx y hy a b ha hb hab, + calc ‖a • x + b • y‖ ≤ ‖a • x‖ + ‖b • y‖ : norm_add_le _ _ + ... = a * ‖x‖ + b * ‖y‖ + : by rw [norm_smul, norm_smul, real.norm_of_nonneg ha, real.norm_of_nonneg hb]⟩ + +/-- The norm on a real normed space is convex on the whole space. See also `seminorm.convex_on` +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) (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) + +lemma convex_on_univ_dist (z : E) : convex_on ℝ univ (λz', dist z' z) := +convex_on_dist z convex_univ + +lemma convex_ball (a : E) (r : ℝ) : convex ℝ (metric.ball a r) := +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 + +/-- 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) : + ∃ x' ∈ s, dist x y ≤ dist x' y := +(convex_on_dist y (convex_convex_hull ℝ _)).exists_ge_of_mem_convex_hull hx + +/-- Given a point `x` in the convex hull of `s` and a point `y` in the convex hull of `t`, +there exist points `x' ∈ s` and `y' ∈ t` at distance at least `dist x y`. -/ +lemma convex_hull_exists_dist_ge2 {s t : set E} {x y : E} + (hx : x ∈ convex_hull ℝ s) (hy : y ∈ convex_hull ℝ t) : + ∃ (x' ∈ s) (y' ∈ t), dist x y ≤ dist x' y' := +begin + rcases convex_hull_exists_dist_ge hx y with ⟨x', hx', Hx'⟩, + rcases convex_hull_exists_dist_ge hy x' with ⟨y', hy', Hy'⟩, + use [x', hx', y', hy'], + exact le_trans Hx' (dist_comm y x' ▸ dist_comm y' x' ▸ Hy') +end + +/-- Emetric diameter of the convex hull of a set `s` equals the emetric diameter of `s. -/ +@[simp] lemma convex_hull_ediam (s : set E) : + emetric.diam (convex_hull ℝ s) = emetric.diam s := +begin + refine (emetric.diam_le $ λ x hx y hy, _).antisymm (emetric.diam_mono $ subset_convex_hull ℝ s), + rcases convex_hull_exists_dist_ge2 hx hy with ⟨x', hx', y', hy', H⟩, + rw edist_dist, + apply le_trans (ennreal.of_real_le_of_real H), + rw ← edist_dist, + exact emetric.edist_le_diam_of_mem hx' hy' +end + +/-- Diameter of the convex hull of a set `s` equals the emetric diameter of `s. -/ +@[simp] lemma convex_hull_diam (s : set E) : + metric.diam (convex_hull ℝ s) = metric.diam s := +by simp only [metric.diam, convex_hull_ediam] + +/-- Convex hull of `s` is bounded if and only if `s` is bounded. -/ +@[simp] lemma bounded_convex_hull {s : set E} : + metric.bounded (convex_hull ℝ s) ↔ metric.bounded s := +by simp only [metric.bounded_iff_ediam_ne_top, convex_hull_ediam] + +@[priority 100] +instance normed_space.path_connected : path_connected_space E := +topological_add_group.path_connected + +@[priority 100] +instance normed_space.loc_path_connected : loc_path_connected_space E := +loc_path_connected_of_bases (λ x, metric.nhds_basis_ball) + (λ x r r_pos, (convex_ball x r).is_path_connected $ by simp [r_pos]) + +lemma dist_add_dist_of_mem_segment {x y z : E} (h : y ∈ [x -[ℝ] z]) : + dist x y + dist y z = dist x z := +begin + simp only [dist_eq_norm, mem_segment_iff_same_ray] at *, + simpa only [sub_add_sub_cancel', norm_sub_rev] using h.norm_add.symm +end + +/-- The set of vectors in the same ray as `x` is connected. -/ +lemma is_connected_set_of_same_ray (x : E) : is_connected {y | same_ray ℝ x y} := +begin + by_cases hx : x = 0, { simpa [hx] using is_connected_univ }, + simp_rw ←exists_nonneg_left_iff_same_ray hx, + exact is_connected_Ici.image _ ((continuous_id.smul continuous_const).continuous_on) +end + +/-- The set of nonzero vectors in the same ray as the nonzero vector `x` is connected. -/ +lemma is_connected_set_of_same_ray_and_ne_zero {x : E} (hx : x ≠ 0) : + is_connected {y | same_ray ℝ x y ∧ y ≠ 0} := +begin + simp_rw ←exists_pos_left_iff_same_ray_and_ne_zero hx, + exact is_connected_Ioi.image _ ((continuous_id.smul continuous_const).continuous_on) +end diff --git a/src/analysis/convex/side.lean b/src/analysis/convex/side.lean index 827cd0dff4c6b..a6941b8d4cc43 100644 --- a/src/analysis/convex/side.lean +++ b/src/analysis/convex/side.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Joseph Myers -/ import analysis.convex.between -import analysis.convex.topology +import analysis.convex.normed import analysis.normed.group.add_torsor /-! diff --git a/src/analysis/convex/strict_convex_space.lean b/src/analysis/convex/strict_convex_space.lean index 39d0309c9d23d..751345c147d10 100644 --- a/src/analysis/convex/strict_convex_space.lean +++ b/src/analysis/convex/strict_convex_space.lean @@ -3,8 +3,8 @@ 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, Yury Kudryashov -/ +import analysis.convex.normed import analysis.convex.strict -import analysis.convex.topology import analysis.normed.order.basic import analysis.normed_space.add_torsor import analysis.normed_space.pointwise diff --git a/src/analysis/convex/topology.lean b/src/analysis/convex/topology.lean index 078d9b74569c5..f36f01234f2f7 100644 --- a/src/analysis/convex/topology.lean +++ b/src/analysis/convex/topology.lean @@ -3,35 +3,25 @@ Copyright (c) 2020 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Alexander Bentkamp, Yury Kudryashov -/ -import analysis.convex.jensen +import analysis.convex.combination import analysis.convex.strict -import analysis.normed.group.pointwise -import topology.algebra.module.finite_dimension -import analysis.normed_space.ray import topology.path_connected import topology.algebra.affine +import topology.algebra.module.basic /-! -# Topological and metric properties of convex sets +# Topological properties of convex sets We prove the following facts: * `convex.interior` : interior of a convex set is convex; * `convex.closure` : closure of a convex set is convex; * `set.finite.compact_convex_hull` : convex hull of a finite set is compact; -* `set.finite.is_closed_convex_hull` : convex hull of a finite set is closed; -* `convex_on_norm`, `convex_on_dist` : norm and distance to a fixed point is convex on any convex - set; -* `convex_on_univ_norm`, `convex_on_univ_dist` : norm and distance to a fixed point is convex on - the whole space; -* `convex_hull_ediam`, `convex_hull_diam` : convex hull of a set has the same (e)metric diameter - as the original set; -* `bounded_convex_hull` : convex hull of a set is bounded if and only if the original set - is bounded. -* `bounded_std_simplex`, `is_closed_std_simplex`, `compact_std_simplex`: topological properties - of the standard simplex; +* `set.finite.is_closed_convex_hull` : convex hull of a finite set is closed. -/ +assert_not_exists has_norm + variables {ι : Type*} {E : Type*} open metric set @@ -53,11 +43,10 @@ lemma std_simplex_subset_closed_ball : std_simplex ℝ ι ⊆ metric.closed_ball 0 1 := begin assume f hf, - rw [metric.mem_closed_ball, dist_zero_right], - refine (nnreal.coe_one ▸ nnreal.coe_le_coe.2 $ finset.sup_le $ λ x hx, _), - change |f x| ≤ 1, - rw [abs_of_nonneg $ hf.1 x], - exact (mem_Icc_of_mem_std_simplex hf x).2 + rw [metric.mem_closed_ball, dist_pi_le_iff zero_le_one], + intros x, + rw [pi.zero_apply, real.dist_0_eq_abs, abs_of_nonneg $ hf.1 x], + exact (mem_Icc_of_mem_std_simplex hf x).2, end variable (ι) @@ -329,120 +318,3 @@ protected lemma topological_add_group.path_connected : path_connected_space E := path_connected_space_iff_univ.mpr $ convex_univ.is_path_connected ⟨(0 : E), trivial⟩ end has_continuous_smul - -/-! ### Normed vector space -/ - -section normed_space -variables [seminormed_add_comm_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 (hs : convex ℝ s) : convex_on ℝ s norm := -⟨hs, λ x hx y hy a b ha hb hab, - calc ‖a • x + b • y‖ ≤ ‖a • x‖ + ‖b • y‖ : norm_add_le _ _ - ... = a * ‖x‖ + b * ‖y‖ - : by rw [norm_smul, norm_smul, real.norm_of_nonneg ha, real.norm_of_nonneg hb]⟩ - -/-- The norm on a real normed space is convex on the whole space. See also `seminorm.convex_on` -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) (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) - -lemma convex_on_univ_dist (z : E) : convex_on ℝ univ (λz', dist z' z) := -convex_on_dist z convex_univ - -lemma convex_ball (a : E) (r : ℝ) : convex ℝ (metric.ball a r) := -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 - -/-- 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) : - ∃ x' ∈ s, dist x y ≤ dist x' y := -(convex_on_dist y (convex_convex_hull ℝ _)).exists_ge_of_mem_convex_hull hx - -/-- Given a point `x` in the convex hull of `s` and a point `y` in the convex hull of `t`, -there exist points `x' ∈ s` and `y' ∈ t` at distance at least `dist x y`. -/ -lemma convex_hull_exists_dist_ge2 {s t : set E} {x y : E} - (hx : x ∈ convex_hull ℝ s) (hy : y ∈ convex_hull ℝ t) : - ∃ (x' ∈ s) (y' ∈ t), dist x y ≤ dist x' y' := -begin - rcases convex_hull_exists_dist_ge hx y with ⟨x', hx', Hx'⟩, - rcases convex_hull_exists_dist_ge hy x' with ⟨y', hy', Hy'⟩, - use [x', hx', y', hy'], - exact le_trans Hx' (dist_comm y x' ▸ dist_comm y' x' ▸ Hy') -end - -/-- Emetric diameter of the convex hull of a set `s` equals the emetric diameter of `s. -/ -@[simp] lemma convex_hull_ediam (s : set E) : - emetric.diam (convex_hull ℝ s) = emetric.diam s := -begin - refine (emetric.diam_le $ λ x hx y hy, _).antisymm (emetric.diam_mono $ subset_convex_hull ℝ s), - rcases convex_hull_exists_dist_ge2 hx hy with ⟨x', hx', y', hy', H⟩, - rw edist_dist, - apply le_trans (ennreal.of_real_le_of_real H), - rw ← edist_dist, - exact emetric.edist_le_diam_of_mem hx' hy' -end - -/-- Diameter of the convex hull of a set `s` equals the emetric diameter of `s. -/ -@[simp] lemma convex_hull_diam (s : set E) : - metric.diam (convex_hull ℝ s) = metric.diam s := -by simp only [metric.diam, convex_hull_ediam] - -/-- Convex hull of `s` is bounded if and only if `s` is bounded. -/ -@[simp] lemma bounded_convex_hull {s : set E} : - metric.bounded (convex_hull ℝ s) ↔ metric.bounded s := -by simp only [metric.bounded_iff_ediam_ne_top, convex_hull_ediam] - -@[priority 100] -instance normed_space.path_connected : path_connected_space E := -topological_add_group.path_connected - -@[priority 100] -instance normed_space.loc_path_connected : loc_path_connected_space E := -loc_path_connected_of_bases (λ x, metric.nhds_basis_ball) - (λ x r r_pos, (convex_ball x r).is_path_connected $ by simp [r_pos]) - -lemma dist_add_dist_of_mem_segment {x y z : E} (h : y ∈ [x -[ℝ] z]) : - dist x y + dist y z = dist x z := -begin - simp only [dist_eq_norm, mem_segment_iff_same_ray] at *, - simpa only [sub_add_sub_cancel', norm_sub_rev] using h.norm_add.symm -end - -/-- The set of vectors in the same ray as `x` is connected. -/ -lemma is_connected_set_of_same_ray (x : E) : is_connected {y | same_ray ℝ x y} := -begin - by_cases hx : x = 0, { simpa [hx] using is_connected_univ }, - simp_rw ←exists_nonneg_left_iff_same_ray hx, - exact is_connected_Ici.image _ ((continuous_id.smul continuous_const).continuous_on) -end - -/-- The set of nonzero vectors in the same ray as the nonzero vector `x` is connected. -/ -lemma is_connected_set_of_same_ray_and_ne_zero {x : E} (hx : x ≠ 0) : - is_connected {y | same_ray ℝ x y ∧ y ≠ 0} := -begin - simp_rw ←exists_pos_left_iff_same_ray_and_ne_zero hx, - exact is_connected_Ioi.image _ ((continuous_id.smul continuous_const).continuous_on) -end - -end normed_space diff --git a/src/analysis/special_functions/bernstein.lean b/src/analysis/special_functions/bernstein.lean index 67df5d5c9e191..a76e8f9207c3e 100644 --- a/src/analysis/special_functions/bernstein.lean +++ b/src/analysis/special_functions/bernstein.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ import algebra.order.field.basic +import analysis.specific_limits.basic import ring_theory.polynomial.bernstein import topology.continuous_function.polynomial import topology.continuous_function.compact diff --git a/src/topology/algebra/module/basic.lean b/src/topology/algebra/module/basic.lean index 014bc026c3e3b..e13655b155398 100644 --- a/src/topology/algebra/module/basic.lean +++ b/src/topology/algebra/module/basic.lean @@ -262,6 +262,27 @@ end end closure +section pi + +lemma linear_map.continuous_on_pi {ι : Type*} {R : Type*} {M : Type*} [finite ι] [semiring R] + [topological_space R] [add_comm_monoid M] [module R M] [topological_space M] + [has_continuous_add M] [has_continuous_smul R M] (f : (ι → R) →ₗ[R] M) : + continuous f := +begin + casesI nonempty_fintype ι, + classical, + -- for the proof, write `f` in the standard basis, and use that each coordinate is a continuous + -- function. + have : (f : (ι → R) → M) = + (λx, ∑ i : ι, x i • (f (λ j, if i = j then 1 else 0))), + by { ext x, exact f.pi_apply_eq_sum_univ x }, + rw this, + refine continuous_finset_sum _ (λi hi, _), + exact (continuous_apply i).smul continuous_const +end + +end pi + /-- Continuous linear maps between modules. We only put the type classes that are necessary for the definition, although in applications `M` and `M₂` will be topological modules over the topological ring `R`. -/ diff --git a/src/topology/algebra/module/finite_dimension.lean b/src/topology/algebra/module/finite_dimension.lean index fece5d836a471..312648ad2a9ca 100644 --- a/src/topology/algebra/module/finite_dimension.lean +++ b/src/topology/algebra/module/finite_dimension.lean @@ -46,29 +46,6 @@ noncomputable theory open set finite_dimensional topological_space filter open_locale big_operators -section semiring - -variables {ι 𝕜 F : Type*} [finite ι] [semiring 𝕜] [topological_space 𝕜] - [add_comm_monoid F] [module 𝕜 F] [topological_space F] - [has_continuous_add F] [has_continuous_smul 𝕜 F] - -/-- A linear map on `ι → 𝕜` (where `ι` is finite) is continuous -/ -lemma linear_map.continuous_on_pi (f : (ι → 𝕜) →ₗ[𝕜] F) : continuous f := -begin - casesI nonempty_fintype ι, - classical, - -- for the proof, write `f` in the standard basis, and use that each coordinate is a continuous - -- function. - have : (f : (ι → 𝕜) → F) = - (λx, ∑ i : ι, x i • (f (λ j, if i = j then 1 else 0))), - by { ext x, exact f.pi_apply_eq_sum_univ x }, - rw this, - refine continuous_finset_sum _ (λi hi, _), - exact (continuous_apply i).smul continuous_const -end - -end semiring - section field variables {𝕜 E F : Type*} [field 𝕜] [topological_space 𝕜] [add_comm_group E] [module 𝕜 E] diff --git a/src/topology/metric_space/metrizable.lean b/src/topology/metric_space/metrizable.lean index d158d8e7691aa..250b0a3467d31 100644 --- a/src/topology/metric_space/metrizable.lean +++ b/src/topology/metric_space/metrizable.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ +import analysis.specific_limits.basic import topology.urysohns_lemma import topology.continuous_function.bounded import topology.uniform_space.cauchy diff --git a/src/topology/tietze_extension.lean b/src/topology/tietze_extension.lean index def8be075ac0c..abaea016c04ac 100644 --- a/src/topology/tietze_extension.lean +++ b/src/topology/tietze_extension.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Yury G. Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury G. Kudryashov -/ +import analysis.specific_limits.basic import data.set.intervals.iso_Ioo import topology.algebra.order.monotone_continuity import topology.urysohns_bounded From 6f9f36364eae3f42368b04858fd66d6d9ae730d8 Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Fri, 20 Jan 2023 23:55:55 +0000 Subject: [PATCH 0310/1291] chore(group_theory/subgroup/basic): split out finiteness (#18242) It feels like it should be possible to define subgroups without relying on all the infrastructure about finite sets and types, and it turns out that it is with fairly limited work. This also has the advantage of removing a few hundred lines of code from what's still one of mathlib's biggest files. --- src/algebra/module/submodule/basic.lean | 2 + src/algebra/periodic.lean | 2 + src/group_theory/commutator.lean | 2 +- src/group_theory/coset.lean | 1 + src/group_theory/free_group.lean | 1 + src/group_theory/group_action/basic.lean | 3 +- src/group_theory/noncomm_pi_coprod.lean | 39 ++++ src/group_theory/perm/subgroup.lean | 2 +- src/group_theory/quotient_group.lean | 1 + src/group_theory/subgroup/basic.lean | 217 +-------------------- src/group_theory/subgroup/finite.lean | 238 +++++++++++++++++++++++ 11 files changed, 293 insertions(+), 215 deletions(-) create mode 100644 src/group_theory/subgroup/finite.lean diff --git a/src/algebra/module/submodule/basic.lean b/src/algebra/module/submodule/basic.lean index 1ffe24a81fd9e..c476f739d5f20 100644 --- a/src/algebra/module/submodule/basic.lean +++ b/src/algebra/module/submodule/basic.lean @@ -6,6 +6,8 @@ Authors: Nathaniel Thomas, Jeremy Avigad, Johannes Hölzl, Mario Carneiro import algebra.module.linear_map import algebra.module.equiv import group_theory.group_action.sub_mul_action +import group_theory.submonoid.membership + /-! # Submodules of a module diff --git a/src/algebra/periodic.lean b/src/algebra/periodic.lean index 9404f8a061254..097404f758d10 100644 --- a/src/algebra/periodic.lean +++ b/src/algebra/periodic.lean @@ -3,12 +3,14 @@ Copyright (c) 2021 Benjamin Davidson. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Benjamin Davidson -/ +import algebra.big_operators.basic import algebra.field.opposite import algebra.module.basic import algebra.order.archimedean import data.int.parity import group_theory.coset import group_theory.subgroup.zpowers +import group_theory.submonoid.membership /-! # Periodicity diff --git a/src/group_theory/commutator.lean b/src/group_theory/commutator.lean index a70850a6c1604..3206dab1db37a 100644 --- a/src/group_theory/commutator.lean +++ b/src/group_theory/commutator.lean @@ -5,7 +5,7 @@ Authors: Jordan Brown, Thomas Browning, Patrick Lutz -/ import data.bracket -import group_theory.subgroup.basic +import group_theory.subgroup.finite import tactic.group /-! diff --git a/src/group_theory/coset.lean b/src/group_theory/coset.lean index 2eee06e3571d5..69843a35df0c7 100644 --- a/src/group_theory/coset.lean +++ b/src/group_theory/coset.lean @@ -5,6 +5,7 @@ Authors: Mitchell Rowett, Scott Morrison -/ import algebra.quotient +import data.fintype.prod import group_theory.group_action.basic import group_theory.subgroup.mul_opposite import tactic.group diff --git a/src/group_theory/free_group.lean b/src/group_theory/free_group.lean index 8c477a913ba27..ed5d485a08040 100644 --- a/src/group_theory/free_group.lean +++ b/src/group_theory/free_group.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau -/ import data.fintype.basic +import data.list.sublists import group_theory.subgroup.basic /-! diff --git a/src/group_theory/group_action/basic.lean b/src/group_theory/group_action/basic.lean index d3e5fd0d59cb6..faec650307fa1 100644 --- a/src/group_theory/group_action/basic.lean +++ b/src/group_theory/group_action/basic.lean @@ -3,6 +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 -/ +import data.fintype.card import group_theory.group_action.defs import group_theory.group_action.group import data.setoid.basic @@ -28,7 +29,7 @@ of `•` belong elsewhere. universes u v w variables {α : Type u} {β : Type v} {γ : Type w} -open_locale big_operators pointwise +open_locale pointwise open function namespace mul_action diff --git a/src/group_theory/noncomm_pi_coprod.lean b/src/group_theory/noncomm_pi_coprod.lean index 500d7f5ad3f3e..3e2db788d60d9 100644 --- a/src/group_theory/noncomm_pi_coprod.lean +++ b/src/group_theory/noncomm_pi_coprod.lean @@ -7,6 +7,7 @@ import group_theory.order_of_element import data.finset.noncomm_prod import data.fintype.big_operators import data.nat.gcd.big_operators +import order.sup_indep /-! # Canonical homomorphism from a finite family of monoids @@ -42,6 +43,44 @@ images of different morphisms commute, we obtain a canonical morphism open_locale big_operators +namespace subgroup + +variables {G : Type*} [group G] + +/-- `finset.noncomm_prod` is “injective” in `f` if `f` maps into independent subgroups. This +generalizes (one direction of) `subgroup.disjoint_iff_mul_eq_one`. -/ +@[to_additive "`finset.noncomm_sum` is “injective” in `f` if `f` maps into independent subgroups. +This generalizes (one direction of) `add_subgroup.disjoint_iff_add_eq_zero`. "] +lemma eq_one_of_noncomm_prod_eq_one_of_independent {ι : Type*} (s : finset ι) (f : ι → G) (comm) + (K : ι → subgroup G) (hind : complete_lattice.independent K) (hmem : ∀ (x ∈ s), f x ∈ K x) + (heq1 : s.noncomm_prod f comm = 1) : ∀ (i ∈ s), f i = 1 := +begin + classical, + revert heq1, + induction s using finset.induction_on with i s hnmem ih, + { simp, }, + { have hcomm := comm.mono (finset.coe_subset.2 $ finset.subset_insert _ _), + simp only [finset.forall_mem_insert] at hmem, + have hmem_bsupr: s.noncomm_prod f hcomm ∈ ⨆ (i ∈ (s : set ι)), K i, + { refine subgroup.noncomm_prod_mem _ _ _, + intros x hx, + have : K x ≤ ⨆ (i ∈ (s : set ι)), K i := le_supr₂ x hx, + exact this (hmem.2 x hx), }, + intro heq1, + rw finset.noncomm_prod_insert_of_not_mem _ _ _ _ hnmem at heq1, + have hnmem' : i ∉ (s : set ι), by simpa, + obtain ⟨heq1i : f i = 1, heq1S : s.noncomm_prod f _ = 1⟩ := + subgroup.disjoint_iff_mul_eq_one.mp (hind.disjoint_bsupr hnmem') hmem.1 hmem_bsupr heq1, + intros i h, + simp only [finset.mem_insert] at h, + rcases h with ⟨rfl | _⟩, + { exact heq1i }, + { exact ih hcomm hmem.2 heq1S _ h } } +end + +end subgroup + + section family_of_monoids variables {M : Type*} [monoid M] diff --git a/src/group_theory/perm/subgroup.lean b/src/group_theory/perm/subgroup.lean index 6bfaa6fafde50..cf8f21c6a40c3 100644 --- a/src/group_theory/perm/subgroup.lean +++ b/src/group_theory/perm/subgroup.lean @@ -5,7 +5,7 @@ Authors: Eric Wieser -/ import group_theory.perm.basic import data.fintype.perm -import group_theory.subgroup.basic +import group_theory.subgroup.finite /-! # Lemmas about subgroups within the permutations (self-equivalences) of a type `α` diff --git a/src/group_theory/quotient_group.lean b/src/group_theory/quotient_group.lean index 14fc5da61645d..7d80d10bef0c1 100644 --- a/src/group_theory/quotient_group.lean +++ b/src/group_theory/quotient_group.lean @@ -7,6 +7,7 @@ This file is to a certain extent based on `quotient_module.lean` by Johannes Hö -/ import group_theory.congruence import group_theory.coset +import group_theory.subgroup.finite import group_theory.subgroup.pointwise /-! diff --git a/src/group_theory/subgroup/basic.lean b/src/group_theory/subgroup/basic.lean index c84d3f099aae4..8e4c2f27970bd 100644 --- a/src/group_theory/subgroup/basic.lean +++ b/src/group_theory/subgroup/basic.lean @@ -7,12 +7,10 @@ import algebra.group.conj import algebra.module.basic import algebra.order.group.inj_surj import data.countable.basic -import data.set.finite import group_theory.submonoid.centralizer -import group_theory.submonoid.membership import logic.encodable.basic import order.atoms -import order.sup_indep +import tactic.apply_fun /-! # Subgroups @@ -86,7 +84,6 @@ subgroup, subgroups -/ open function -open_locale big_operators variables {G G' : Type*} [group G] [group G'] variables {A : Type*} [add_group A] @@ -324,14 +321,6 @@ lemma coe_to_submonoid (K : subgroup G) : (K.to_submonoid : set G) = K := rfl @[simp, to_additive] lemma mem_to_submonoid (K : subgroup G) (x : G) : x ∈ K.to_submonoid ↔ x ∈ K := iff.rfl -@[to_additive] -instance (K : subgroup G) [d : decidable_pred (∈ K)] [fintype G] : fintype K := -show fintype {g : G // g ∈ K}, from infer_instance - -@[to_additive] -instance (K : subgroup G) [finite G] : finite K := -subtype.finite - @[to_additive] theorem to_submonoid_injective : function.injective (to_submonoid : subgroup G → submonoid G) := @@ -457,37 +446,6 @@ mul_mem_cancel_right h @[to_additive] protected lemma mul_mem_cancel_left {x y : G} (h : x ∈ H) : x * y ∈ H ↔ y ∈ H := mul_mem_cancel_left h -/-- Product of a list of elements in a subgroup is in the subgroup. -/ -@[to_additive "Sum of a list of elements in an `add_subgroup` is in the `add_subgroup`."] -protected lemma list_prod_mem {l : list G} : (∀ x ∈ l, x ∈ K) → l.prod ∈ K := -list_prod_mem - -/-- Product of a multiset of elements in a subgroup of a `comm_group` is in the subgroup. -/ -@[to_additive "Sum of a multiset of elements in an `add_subgroup` of an `add_comm_group` -is in the `add_subgroup`."] -protected lemma multiset_prod_mem {G} [comm_group G] (K : subgroup G) (g : multiset G) : - (∀ a ∈ g, a ∈ K) → g.prod ∈ K := multiset_prod_mem g - -@[to_additive] -lemma multiset_noncomm_prod_mem (K : subgroup G) (g : multiset G) (comm) : - (∀ a ∈ g, a ∈ K) → g.noncomm_prod comm ∈ K := -K.to_submonoid.multiset_noncomm_prod_mem g comm - -/-- Product of elements of a subgroup of a `comm_group` indexed by a `finset` is in the - subgroup. -/ -@[to_additive "Sum of elements in an `add_subgroup` of an `add_comm_group` indexed by a `finset` -is in the `add_subgroup`."] -protected lemma prod_mem {G : Type*} [comm_group G] (K : subgroup G) - {ι : Type*} {t : finset ι} {f : ι → G} (h : ∀ c ∈ t, f c ∈ K) : - ∏ c in t, f c ∈ K := -prod_mem h - -@[to_additive] -lemma noncomm_prod_mem (K : subgroup G) {ι : Type*} {t : finset ι} {f : ι → G} (comm) : - (∀ c ∈ t, f c ∈ K) → t.noncomm_prod f comm ∈ K := -K.to_submonoid.noncomm_prod_mem t f comm - - @[to_additive add_subgroup.nsmul_mem] protected lemma pow_mem {x : G} (hx : x ∈ K) : ∀ n : ℕ, x ^ n ∈ K := pow_mem hx @@ -581,19 +539,6 @@ def subtype : H →* G := ⟨coe, rfl, λ _ _, rfl⟩ @[to_additive] lemma subtype_injective : function.injective (subtype H) := subtype.coe_injective -@[simp, norm_cast, to_additive] theorem coe_list_prod (l : list H) : - (l.prod : G) = (l.map coe).prod := -submonoid_class.coe_list_prod l - -@[simp, norm_cast, to_additive] theorem coe_multiset_prod {G} [comm_group G] (H : subgroup G) - (m : multiset H) : (m.prod : G) = (m.map coe).prod := -submonoid_class.coe_multiset_prod m - -@[simp, norm_cast, to_additive] theorem coe_finset_prod {ι G} [comm_group G] (H : subgroup G) - (f : ι → H) (s : finset ι) : - ↑(∏ i in s, f i) = (∏ i in s, f i : G) := -submonoid_class.coe_finset_prod f s - /-- The inclusion homomorphism from a subgroup `H` contained in `K` to `K`. -/ @[to_additive "The inclusion homomorphism from a additive subgroup `H` contained in `K` to `K`."] def inclusion {H K : subgroup G} (h : H ≤ K) : H →* K := @@ -664,35 +609,6 @@ end ⟨λ ⟨g, hg⟩, by { haveI : subsingleton (H : set G) := by { rw hg, apply_instance }, exact H.eq_bot_of_subsingleton }, λ h, ⟨1, set_like.ext'_iff.mp h⟩⟩ -@[to_additive] instance fintype_bot : fintype (⊥ : subgroup G) := ⟨{1}, -by {rintro ⟨x, ⟨hx⟩⟩, exact finset.mem_singleton_self _}⟩ - -/- curly brackets `{}` are used here instead of instance brackets `[]` because - the instance in a goal is often not the same as the one inferred by type class inference. -/ -@[simp, to_additive] lemma card_bot {_ : fintype ↥(⊥ : subgroup G)} : - fintype.card (⊥ : subgroup G) = 1 := -fintype.card_eq_one_iff.2 - ⟨⟨(1 : G), set.mem_singleton 1⟩, λ ⟨y, hy⟩, subtype.eq $ subgroup.mem_bot.1 hy⟩ - -@[to_additive] lemma eq_top_of_card_eq [fintype H] [fintype G] - (h : fintype.card H = fintype.card G) : H = ⊤ := -begin - haveI : fintype (H : set G) := ‹fintype H›, - rw [set_like.ext'_iff, coe_top, ← finset.coe_univ, ← (H : set G).coe_to_finset, finset.coe_inj, - ← finset.card_eq_iff_eq_univ, ← h, set.to_finset_card], - congr -end - -@[to_additive] lemma eq_top_of_le_card [fintype H] [fintype G] - (h : fintype.card G ≤ fintype.card H) : H = ⊤ := -eq_top_of_card_eq H (le_antisymm (fintype.card_le_of_injective coe subtype.coe_injective) h) - -@[to_additive] lemma eq_bot_of_card_le [fintype H] (h : fintype.card H ≤ 1) : H = ⊥ := -let _ := fintype.card_le_one_iff_subsingleton.mp h in by exactI eq_bot_of_subsingleton H - -@[to_additive] lemma eq_bot_of_card_eq [fintype H] (h : fintype.card H = 1) : H = ⊥ := -H.eq_bot_of_card_le (le_of_eq h) - @[to_additive] lemma nontrivial_iff_exists_ne_one (H : subgroup G) : nontrivial H ↔ ∃ x ∈ H, x ≠ (1:G) := subtype.nontrivial_iff_exists_ne (λ x, x ∈ H) (1 : H) @@ -718,14 +634,6 @@ begin rw nontrivial_iff_exists_ne_one end -@[to_additive] lemma card_le_one_iff_eq_bot [fintype H] : fintype.card H ≤ 1 ↔ H = ⊥ := -⟨λ h, (eq_bot_iff_forall _).2 - (λ x hx, by simpa [subtype.ext_iff] using fintype.card_le_one_iff.1 h ⟨x, hx⟩ 1), - λ h, by simp [h]⟩ - -@[to_additive] lemma one_lt_card_iff_ne_bot [fintype H] : 1 < fintype.card H ↔ H ≠ ⊥ := -lt_iff_not_le.trans H.card_le_one_iff_eq_bot.not - /-- The inf of two subgroups is their intersection. -/ @[to_additive "The inf of two `add_subgroups`s is their intersection."] instance : has_inf (subgroup G) := @@ -1382,51 +1290,6 @@ begin { simp [heq, one_mem], }, } end -@[to_additive] -lemma pi_mem_of_mul_single_mem_aux [decidable_eq η] (I : finset η) {H : subgroup (Π i, f i) } - (x : Π i, f i) (h1 : ∀ i, i ∉ I → x i = 1) (h2 : ∀ i, i ∈ I → pi.mul_single i (x i) ∈ H ) : - x ∈ H := -begin - induction I using finset.induction_on with i I hnmem ih generalizing x, - { convert one_mem H, - ext i, - exact (h1 i (not_mem_empty i)) }, - { have : x = function.update x i 1 * pi.mul_single i (x i), - { ext j, - by_cases heq : j = i, - { subst heq, simp, }, - { simp [heq], }, }, - rw this, clear this, - apply mul_mem, - { apply ih; clear ih, - { intros j hj, - by_cases heq : j = i, - { subst heq, simp, }, - { simp [heq], apply h1 j, simpa [heq] using hj, } }, - { intros j hj, - have : j ≠ i, by { rintro rfl, contradiction }, - simp [this], - exact h2 _ (finset.mem_insert_of_mem hj), }, }, - { apply h2, simp, } } -end - -@[to_additive] -lemma pi_mem_of_mul_single_mem [finite η] [decidable_eq η] {H : subgroup (Π i, f i)} - (x : Π i, f i) (h : ∀ i, pi.mul_single i (x i) ∈ H) : x ∈ H := -by { casesI nonempty_fintype η, - exact pi_mem_of_mul_single_mem_aux finset.univ x (by simp) (λ i _, h i) } - -/-- For finite index types, the `subgroup.pi` is generated by the embeddings of the groups. -/ -@[to_additive "For finite index types, the `subgroup.pi` is generated by the embeddings of the -additive groups."] -lemma pi_le_iff [decidable_eq η] [finite η] {H : Π i, subgroup (f i)} {J : subgroup (Π i, f i)} : - pi univ H ≤ J ↔ ∀ i : η, map (monoid_hom.single f i) (H i) ≤ J := -begin - split, - { rintros h i _ ⟨x, hx, rfl⟩, apply h, simpa using hx }, - { exact λ h x hx, pi_mem_of_mul_single_mem x (λ i, h i (mem_map_of_mem _ (hx i trivial))), } -end - @[to_additive] lemma pi_eq_bot_iff (H : Π i, subgroup (f i)) : pi set.univ H = ⊥ ↔ ∀ i, H i = ⊥ := @@ -1576,8 +1439,8 @@ variable {G} @[to_additive] lemma mem_center_iff {z : G} : z ∈ center G ↔ ∀ g, g * z = z * g := iff.rfl -instance decidable_mem_center [decidable_eq G] [fintype G] : decidable_pred (∈ center G) := -λ _, decidable_of_iff' _ mem_center_iff +instance decidable_mem_center (z : G) [decidable (∀ g, g * z = z * g)] : decidable (z ∈ center G) := +decidable_of_iff' _ mem_center_iff @[to_additive] instance center_characteristic : (center G).characteristic := begin @@ -1621,16 +1484,6 @@ def set_normalizer (S : set G) : subgroup G := inv_mem' := λ a (ha : ∀ n, n ∈ S ↔ a * n * a⁻¹ ∈ S) n, by { rw [ha (a⁻¹ * n * a⁻¹⁻¹)], simp [mul_assoc] } } -lemma mem_normalizer_fintype {S : set G} [finite S] {x : G} - (h : ∀ n, n ∈ S → x * n * x⁻¹ ∈ S) : x ∈ subgroup.set_normalizer S := -by haveI := classical.prop_decidable; casesI nonempty_fintype S; -haveI := set.fintype_image S (λ n, x * n * x⁻¹); exact -λ n, ⟨h n, λ h₁, -have heq : (λ n, x * n * x⁻¹) '' S = S := set.eq_of_subset_of_card_le - (λ n ⟨y, hy⟩, hy.2 ▸ h y hy.1) (by rw set.card_image_of_injective S conj_injective), -have x * n * x⁻¹ ∈ (λ n, x * n * x⁻¹) '' S := heq.symm ▸ h₁, -let ⟨y, hy⟩ := this in conj_injective hy.2 ▸ hy.1⟩ - variable {H} @[to_additive] lemma mem_normalizer_iff {g : G} : g ∈ H.normalizer ↔ ∀ h, h ∈ H ↔ g * h * g⁻¹ ∈ H := @@ -1960,11 +1813,6 @@ open subgroup def range (f : G →* N) : subgroup N := subgroup.copy ((⊤ : subgroup G).map f) (set.range f) (by simp [set.ext_iff]) -@[to_additive] -instance decidable_mem_range (f : G →* N) [fintype G] [decidable_eq N] : - decidable_pred (∈ f.range) := -λ x, fintype.decidable_exists_fintype - @[simp, to_additive] lemma coe_range (f : G →* N) : (f.range : set N) = set.range f := rfl @@ -2198,32 +2046,6 @@ lemma map_closure (f : G →* N) (s : set G) : set.image_preimage.l_comm_of_u_comm (subgroup.gc_map_comap f) (subgroup.gi N).gc (subgroup.gi G).gc (λ t, rfl) --- this instance can't go just after the definition of `mrange` because `fintype` is --- not imported at that stage - -/-- The range of a finite monoid under a monoid homomorphism is finite. -Note: this instance can form a diamond with `subtype.fintype` in the -presence of `fintype N`. -/ -@[to_additive "The range of a finite additive monoid under an additive monoid homomorphism is -finite. - -Note: this instance can form a diamond with `subtype.fintype` or `subgroup.fintype` in the -presence of `fintype N`."] -instance fintype_mrange {M N : Type*} [monoid M] [monoid N] [fintype M] [decidable_eq N] - (f : M →* N) : fintype (mrange f) := -set.fintype_range f - -/-- The range of a finite group under a group homomorphism is finite. - -Note: this instance can form a diamond with `subtype.fintype` or `subgroup.fintype` in the -presence of `fintype N`. -/ -@[to_additive "The range of a finite additive group under an additive group homomorphism is finite. - -Note: this instance can form a diamond with `subtype.fintype` or `subgroup.fintype` in the -presence of `fintype N`."] -instance fintype_range [fintype G] [decidable_eq N] (f : G →* N) : fintype (range f) := -set.fintype_range f - end monoid_hom namespace subgroup @@ -2812,37 +2634,6 @@ begin ←subtype.ext_iff, ←subtype.ext_iff, eq_comm, ←prod.ext_iff] at hxy, end -/-- `finset.noncomm_prod` is “injective” in `f` if `f` maps into independent subgroups. This -generalizes (one direction of) `subgroup.disjoint_iff_mul_eq_one`. -/ -@[to_additive "`finset.noncomm_sum` is “injective” in `f` if `f` maps into independent subgroups. -This generalizes (one direction of) `add_subgroup.disjoint_iff_add_eq_zero`. "] -lemma eq_one_of_noncomm_prod_eq_one_of_independent {ι : Type*} (s : finset ι) (f : ι → G) (comm) - (K : ι → subgroup G) (hind : complete_lattice.independent K) (hmem : ∀ (x ∈ s), f x ∈ K x) - (heq1 : s.noncomm_prod f comm = 1) : ∀ (i ∈ s), f i = 1 := -begin - classical, - revert heq1, - induction s using finset.induction_on with i s hnmem ih, - { simp, }, - { have hcomm := comm.mono (finset.coe_subset.2 $ finset.subset_insert _ _), - simp only [finset.forall_mem_insert] at hmem, - have hmem_bsupr: s.noncomm_prod f hcomm ∈ ⨆ (i ∈ (s : set ι)), K i, - { refine subgroup.noncomm_prod_mem _ _ _, - intros x hx, - have : K x ≤ ⨆ (i ∈ (s : set ι)), K i := le_supr₂ x hx, - exact this (hmem.2 x hx), }, - intro heq1, - rw finset.noncomm_prod_insert_of_not_mem _ _ _ _ hnmem at heq1, - have hnmem' : i ∉ (s : set ι), by simpa, - obtain ⟨heq1i : f i = 1, heq1S : s.noncomm_prod f _ = 1⟩ := - subgroup.disjoint_iff_mul_eq_one.mp (hind.disjoint_bsupr hnmem') hmem.1 hmem_bsupr heq1, - intros i h, - simp only [finset.mem_insert] at h, - rcases h with ⟨rfl | _⟩, - { exact heq1i }, - { exact ih hcomm hmem.2 heq1S _ h } } -end - end subgroup namespace is_conj @@ -2875,3 +2666,5 @@ begin end end is_conj + +assert_not_exists multiset diff --git a/src/group_theory/subgroup/finite.lean b/src/group_theory/subgroup/finite.lean new file mode 100644 index 0000000000000..b8de0c882f0db --- /dev/null +++ b/src/group_theory/subgroup/finite.lean @@ -0,0 +1,238 @@ +/- +Copyright (c) 2020 Kexing Ying. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Kexing Ying +-/ + +import data.set.finite +import group_theory.subgroup.basic +import group_theory.submonoid.membership + +/-! +# Subgroups + +This file provides some result on multiplicative and additive subgroups in the finite context. + +## Tags +subgroup, subgroups +-/ + +open_locale big_operators + +variables {G : Type*} [group G] +variables {A : Type*} [add_group A] + +namespace subgroup + +@[to_additive] +instance (K : subgroup G) [d : decidable_pred (∈ K)] [fintype G] : fintype K := +show fintype {g : G // g ∈ K}, from infer_instance + +@[to_additive] +instance (K : subgroup G) [finite G] : finite K := +subtype.finite + +end subgroup + +/-! +### Conversion to/from `additive`/`multiplicative` +-/ +namespace subgroup + +variables (H K : subgroup G) + +/-- Product of a list of elements in a subgroup is in the subgroup. -/ +@[to_additive "Sum of a list of elements in an `add_subgroup` is in the `add_subgroup`."] +protected lemma list_prod_mem {l : list G} : (∀ x ∈ l, x ∈ K) → l.prod ∈ K := +list_prod_mem + +/-- Product of a multiset of elements in a subgroup of a `comm_group` is in the subgroup. -/ +@[to_additive "Sum of a multiset of elements in an `add_subgroup` of an `add_comm_group` +is in the `add_subgroup`."] +protected lemma multiset_prod_mem {G} [comm_group G] (K : subgroup G) (g : multiset G) : + (∀ a ∈ g, a ∈ K) → g.prod ∈ K := multiset_prod_mem g + +@[to_additive] +lemma multiset_noncomm_prod_mem (K : subgroup G) (g : multiset G) (comm) : + (∀ a ∈ g, a ∈ K) → g.noncomm_prod comm ∈ K := +K.to_submonoid.multiset_noncomm_prod_mem g comm + +/-- Product of elements of a subgroup of a `comm_group` indexed by a `finset` is in the + subgroup. -/ +@[to_additive "Sum of elements in an `add_subgroup` of an `add_comm_group` indexed by a `finset` +is in the `add_subgroup`."] +protected lemma prod_mem {G : Type*} [comm_group G] (K : subgroup G) + {ι : Type*} {t : finset ι} {f : ι → G} (h : ∀ c ∈ t, f c ∈ K) : + ∏ c in t, f c ∈ K := +prod_mem h + +@[to_additive] +lemma noncomm_prod_mem (K : subgroup G) {ι : Type*} {t : finset ι} {f : ι → G} (comm) : + (∀ c ∈ t, f c ∈ K) → t.noncomm_prod f comm ∈ K := +K.to_submonoid.noncomm_prod_mem t f comm + +@[simp, norm_cast, to_additive] theorem coe_list_prod (l : list H) : + (l.prod : G) = (l.map coe).prod := +submonoid_class.coe_list_prod l + +@[simp, norm_cast, to_additive] theorem coe_multiset_prod {G} [comm_group G] (H : subgroup G) + (m : multiset H) : (m.prod : G) = (m.map coe).prod := +submonoid_class.coe_multiset_prod m + +@[simp, norm_cast, to_additive] theorem coe_finset_prod {ι G} [comm_group G] (H : subgroup G) + (f : ι → H) (s : finset ι) : + ↑(∏ i in s, f i) = (∏ i in s, f i : G) := +submonoid_class.coe_finset_prod f s + +@[to_additive] instance fintype_bot : fintype (⊥ : subgroup G) := ⟨{1}, +by {rintro ⟨x, ⟨hx⟩⟩, exact finset.mem_singleton_self _}⟩ + +/- curly brackets `{}` are used here instead of instance brackets `[]` because + the instance in a goal is often not the same as the one inferred by type class inference. -/ +@[simp, to_additive] lemma card_bot {_ : fintype ↥(⊥ : subgroup G)} : + fintype.card (⊥ : subgroup G) = 1 := +fintype.card_eq_one_iff.2 + ⟨⟨(1 : G), set.mem_singleton 1⟩, λ ⟨y, hy⟩, subtype.eq $ subgroup.mem_bot.1 hy⟩ + +@[to_additive] lemma eq_top_of_card_eq [fintype H] [fintype G] + (h : fintype.card H = fintype.card G) : H = ⊤ := +begin + haveI : fintype (H : set G) := ‹fintype H›, + rw [set_like.ext'_iff, coe_top, ← finset.coe_univ, ← (H : set G).coe_to_finset, finset.coe_inj, + ← finset.card_eq_iff_eq_univ, ← h, set.to_finset_card], + congr +end + +@[to_additive] lemma eq_top_of_le_card [fintype H] [fintype G] + (h : fintype.card G ≤ fintype.card H) : H = ⊤ := +eq_top_of_card_eq H (le_antisymm (fintype.card_le_of_injective coe subtype.coe_injective) h) + +@[to_additive] lemma eq_bot_of_card_le [fintype H] (h : fintype.card H ≤ 1) : H = ⊥ := +let _ := fintype.card_le_one_iff_subsingleton.mp h in by exactI eq_bot_of_subsingleton H + +@[to_additive] lemma eq_bot_of_card_eq [fintype H] (h : fintype.card H = 1) : H = ⊥ := +H.eq_bot_of_card_le (le_of_eq h) + +@[to_additive] lemma card_le_one_iff_eq_bot [fintype H] : fintype.card H ≤ 1 ↔ H = ⊥ := +⟨λ h, (eq_bot_iff_forall _).2 + (λ x hx, by simpa [subtype.ext_iff] using fintype.card_le_one_iff.1 h ⟨x, hx⟩ 1), + λ h, by simp [h]⟩ + +@[to_additive] lemma one_lt_card_iff_ne_bot [fintype H] : 1 < fintype.card H ↔ H ≠ ⊥ := +lt_iff_not_le.trans H.card_le_one_iff_eq_bot.not + +end subgroup + +namespace subgroup + +section pi + +open set + +variables {η : Type*} {f : η → Type*} [∀ i, group (f i)] + +@[to_additive] +lemma pi_mem_of_mul_single_mem_aux [decidable_eq η] (I : finset η) {H : subgroup (Π i, f i) } + (x : Π i, f i) (h1 : ∀ i, i ∉ I → x i = 1) (h2 : ∀ i, i ∈ I → pi.mul_single i (x i) ∈ H ) : + x ∈ H := +begin + induction I using finset.induction_on with i I hnmem ih generalizing x, + { convert one_mem H, + ext i, + exact (h1 i (not_mem_empty i)) }, + { have : x = function.update x i 1 * pi.mul_single i (x i), + { ext j, + by_cases heq : j = i, + { subst heq, simp, }, + { simp [heq], }, }, + rw this, clear this, + apply mul_mem, + { apply ih; clear ih, + { intros j hj, + by_cases heq : j = i, + { subst heq, simp, }, + { simp [heq], apply h1 j, simpa [heq] using hj, } }, + { intros j hj, + have : j ≠ i, by { rintro rfl, contradiction }, + simp [this], + exact h2 _ (finset.mem_insert_of_mem hj), }, }, + { apply h2, simp, } } +end + +@[to_additive] +lemma pi_mem_of_mul_single_mem [finite η] [decidable_eq η] {H : subgroup (Π i, f i)} + (x : Π i, f i) (h : ∀ i, pi.mul_single i (x i) ∈ H) : x ∈ H := +by { casesI nonempty_fintype η, + exact pi_mem_of_mul_single_mem_aux finset.univ x (by simp) (λ i _, h i) } + +/-- For finite index types, the `subgroup.pi` is generated by the embeddings of the groups. -/ +@[to_additive "For finite index types, the `subgroup.pi` is generated by the embeddings of the +additive groups."] +lemma pi_le_iff [decidable_eq η] [finite η] {H : Π i, subgroup (f i)} {J : subgroup (Π i, f i)} : + pi univ H ≤ J ↔ ∀ i : η, map (monoid_hom.single f i) (H i) ≤ J := +begin + split, + { rintros h i _ ⟨x, hx, rfl⟩, apply h, simpa using hx }, + { exact λ h x hx, pi_mem_of_mul_single_mem x (λ i, h i (mem_map_of_mem _ (hx i trivial))), } +end + +end pi + +end subgroup + +namespace subgroup + +section normalizer + +lemma mem_normalizer_fintype {S : set G} [finite S] {x : G} + (h : ∀ n, n ∈ S → x * n * x⁻¹ ∈ S) : x ∈ subgroup.set_normalizer S := +by haveI := classical.prop_decidable; casesI nonempty_fintype S; +haveI := set.fintype_image S (λ n, x * n * x⁻¹); exact +λ n, ⟨h n, λ h₁, +have heq : (λ n, x * n * x⁻¹) '' S = S := set.eq_of_subset_of_card_le + (λ n ⟨y, hy⟩, hy.2 ▸ h y hy.1) (by rw set.card_image_of_injective S conj_injective), +have x * n * x⁻¹ ∈ (λ n, x * n * x⁻¹) '' S := heq.symm ▸ h₁, +let ⟨y, hy⟩ := this in conj_injective hy.2 ▸ hy.1⟩ + +end normalizer + +end subgroup + +namespace monoid_hom + +variables {N : Type*} [group N] + +open subgroup + +@[to_additive] +instance decidable_mem_range (f : G →* N) [fintype G] [decidable_eq N] : + decidable_pred (∈ f.range) := +λ x, fintype.decidable_exists_fintype + +-- this instance can't go just after the definition of `mrange` because `fintype` is +-- not imported at that stage + +/-- The range of a finite monoid under a monoid homomorphism is finite. +Note: this instance can form a diamond with `subtype.fintype` in the +presence of `fintype N`. -/ +@[to_additive "The range of a finite additive monoid under an additive monoid homomorphism is +finite. + +Note: this instance can form a diamond with `subtype.fintype` or `subgroup.fintype` in the +presence of `fintype N`."] +instance fintype_mrange {M N : Type*} [monoid M] [monoid N] [fintype M] [decidable_eq N] + (f : M →* N) : fintype (mrange f) := +set.fintype_range f + +/-- The range of a finite group under a group homomorphism is finite. + +Note: this instance can form a diamond with `subtype.fintype` or `subgroup.fintype` in the +presence of `fintype N`. -/ +@[to_additive "The range of a finite additive group under an additive group homomorphism is finite. + +Note: this instance can form a diamond with `subtype.fintype` or `subgroup.fintype` in the +presence of `fintype N`."] +instance fintype_range [fintype G] [decidable_eq N] (f : G →* N) : fintype (range f) := +set.fintype_range f + +end monoid_hom From 38fe3f3c8ce2c93f77eb356805f89e99593df483 Mon Sep 17 00:00:00 2001 From: Christopher Hoskin Date: Sat, 21 Jan 2023 10:44:32 +0000 Subject: [PATCH 0311/1291] feat(topology/order/lower_topology): Introduce the lower topology on a preorder (#17426) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR introduces the lower topology on a preorder. It is shown that the lower topology on a partial order is T₀ and the non-empty complements of the upper closures of finite subsets form a basis. It is also shown that the lower topology on the product of two `order_bot` spaces coincides with the product topology of the lower topologies on the component spaces. This is used to show that the inf map `(a,b) → a ⊓ b` on a complete lattice is continuous. The motivation for introducing the lower topology to mathlib is that, along with the Scott Topology (`topology.omega_complete_partial_order`) it is used to define the Lawson topology (Gierz et al, Chapter III). --- src/order/hom/complete_lattice.lean | 18 +- src/order/upper_lower.lean | 10 +- src/topology/order/lower_topology.lean | 230 +++++++++++++++++++++++++ 3 files changed, 245 insertions(+), 13 deletions(-) create mode 100644 src/topology/order/lower_topology.lean diff --git a/src/order/hom/complete_lattice.lean b/src/order/hom/complete_lattice.lean index e37431e776b62..6f1a44ed63de6 100644 --- a/src/order/hom/complete_lattice.lean +++ b/src/order/hom/complete_lattice.lean @@ -627,17 +627,17 @@ See also `complete_lattice_hom.set_preimage`. -/ map_rel_iff' := λ s t, ⟨λ h, by simpa using @monotone_image _ _ e.symm _ _ h, λ h, monotone_image h⟩ } -section +variables [complete_lattice α] (x : α × α) -variables (α) [complete_lattice α] +/-- The map `(a, b) ↦ a ⊔ b` as a `Sup_hom`. -/ +def sup_Sup_hom : Sup_hom (α × α) α := +{ to_fun := λ x, x.1 ⊔ x.2, + map_Sup' := λ s, by simp_rw [prod.fst_Sup, prod.snd_Sup, Sup_image, supr_sup_eq] } /-- The map `(a, b) ↦ a ⊓ b` as an `Inf_hom`. -/ -@[simps] def inf_Inf_hom : Inf_hom (α × α) α := +def inf_Inf_hom : Inf_hom (α × α) α := { to_fun := λ x, x.1 ⊓ x.2, - map_Inf' := λ s, by simp_rw [prod.fst_Inf, prod.snd_Inf, Inf_image, infi_inf_eq], } + map_Inf' := λ s, by simp_rw [prod.fst_Inf, prod.snd_Inf, Inf_image, infi_inf_eq] } -/-- The map `(a, b) ↦ a ⊔ b` as a `Sup_hom`. -/ -@[simps] def sup_Sup_hom : Sup_hom (α × α) α := -{ to_fun := λ x, x.1 ⊔ x.2, map_Sup' := (inf_Inf_hom αᵒᵈ).map_Inf' } - -end +@[simp, norm_cast] lemma sup_Sup_hom_apply : sup_Sup_hom x = x.1 ⊔ x.2 := rfl +@[simp, norm_cast] lemma inf_Inf_hom_apply : inf_Inf_hom x = x.1 ⊓ x.2 := rfl diff --git a/src/order/upper_lower.lean b/src/order/upper_lower.lean index 61b3c5a6c2b40..2bcc1e90c8a6b 100644 --- a/src/order/upper_lower.lean +++ b/src/order/upper_lower.lean @@ -757,13 +757,15 @@ def upper_closure (s : set α) : upper_set α := def lower_closure (s : set α) : lower_set α := ⟨{x | ∃ a ∈ s, x ≤ a}, λ x y h, Exists₂.imp $ λ a _, h.trans⟩ --- We do not tag those two as `simp` to respect the abstraction. -@[norm_cast] lemma coe_upper_closure (s : set α) : ↑(upper_closure s) = {x | ∃ a ∈ s, a ≤ x} := rfl -@[norm_cast] lemma coe_lower_closure (s : set α) : ↑(lower_closure s) = {x | ∃ a ∈ s, x ≤ a} := rfl - @[simp] lemma mem_upper_closure : x ∈ upper_closure s ↔ ∃ a ∈ s, a ≤ x := iff.rfl @[simp] lemma mem_lower_closure : x ∈ lower_closure s ↔ ∃ a ∈ s, x ≤ a := iff.rfl +-- We do not tag those two as `simp` to respect the abstraction. +@[norm_cast] lemma coe_upper_closure (s : set α) : ↑(upper_closure s) = ⋃ a ∈ s, Ici a := +by { ext, simp } +@[norm_cast] lemma coe_lower_closure (s : set α) : ↑(lower_closure s) = ⋃ a ∈ s, Iic a := +by { ext, simp } + lemma subset_upper_closure : s ⊆ upper_closure s := λ x hx, ⟨x, hx, le_rfl⟩ lemma subset_lower_closure : s ⊆ lower_closure s := λ x hx, ⟨x, hx, le_rfl⟩ diff --git a/src/topology/order/lower_topology.lean b/src/topology/order/lower_topology.lean new file mode 100644 index 0000000000000..0412ba2aa4092 --- /dev/null +++ b/src/topology/order/lower_topology.lean @@ -0,0 +1,230 @@ +/- +Copyright (c) 2023 Christopher Hoskin. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Christopher Hoskin +-/ +import topology.order.lattice + +/-! +# Lower topology + +This file introduces the lower topology on a preorder as the topology generated by the complements +of the closed intervals to infinity. + +## Main statements + +- `lower_topology.t0_space` - the lower topology on a partial order is T₀ +- `is_topological_basis.is_topological_basis` - the complements of the upper closures of finite + subsets form a basis for the lower topology +- `lower_topology.to_has_continuous_inf` - the inf map is continuous with respect to the lower + topology + +## Implementation notes + +A type synonym `with_lower_topology` is introduced and for a preorder `α`, `with_lower_topology α` +is made an instance of `topological_space` by the topology generated by the complements of the +closed intervals to infinity. + +We define a mixin class `lower_topology` for the class of types which are both a preorder and a +topology and where the topology is generated by the complements of the closed intervals to infinity. +It is shown that `with_lower_topology α` is an instance of `lower_topology`. + +## Motivation + +The lower topology is used with the `Scott` topology to define the Lawson topology. The restriction +of the lower topology to the spectrum of a complete lattice coincides with the hull-kernel topology. + +## References + +* [Gierz et al, *A Compendium of Continuous Lattices*][GierzEtAl1980] + +## Tags + +lower topology, preorder +-/ + +variables (α β : Type*) + +open set topological_space + +/-- +Type synonym for a preorder equipped with the lower topology +-/ +def with_lower_topology := α + +variables {α β} + +namespace with_lower_topology + +/-- `to_lower` is the identity function to the `with_lower_topology` of a type. -/ +@[pattern] def to_lower : α ≃ with_lower_topology α := equiv.refl _ + +/-- `of_lower` is the identity function from the `with_lower_topology` of a type. -/ +@[pattern] def of_lower : with_lower_topology α ≃ α := equiv.refl _ + +@[simp] lemma to_with_lower_topology_symm_eq : (@to_lower α).symm = of_lower := rfl +@[simp] lemma of_with_lower_topology_symm_eq : (@of_lower α).symm = to_lower := rfl +@[simp] lemma to_lower_of_lower (a : with_lower_topology α) : to_lower (of_lower a) = a := rfl +@[simp] lemma of_lower_to_lower (a : α) : of_lower (to_lower a) = a := rfl +@[simp] lemma to_lower_inj {a b : α} : to_lower a = to_lower b ↔ a = b := iff.rfl +@[simp] lemma of_lower_inj {a b : with_lower_topology α} : of_lower a = of_lower b ↔ a = b := +iff.rfl + +/-- A recursor for `with_lower_topology`. Use as `induction x using with_lower_topology.rec`. -/ +protected def rec {β : with_lower_topology α → Sort*} + (h : Π a, β (to_lower a)) : Π a, β a := λ a, h (of_lower a) + +instance [nonempty α] : nonempty (with_lower_topology α) := ‹nonempty α› +instance [inhabited α] : inhabited (with_lower_topology α) := ‹inhabited α› + +variables [preorder α] + +instance : preorder (with_lower_topology α) := ‹preorder α› + +instance : topological_space (with_lower_topology α) := generate_from {s | ∃ a, (Ici a)ᶜ = s} + +lemma is_open_preimage_of_lower (S : set α) : + is_open (with_lower_topology.of_lower ⁻¹' S) ↔ + (generate_from {s : set α | ∃ (a : α), (Ici a)ᶜ = s}).is_open S := iff.rfl + +lemma is_open_def (T : set (with_lower_topology α)) : + is_open T ↔ (generate_from {s : set α | ∃ (a : α), (Ici a)ᶜ = s}).is_open + (with_lower_topology.to_lower ⁻¹' T) := iff.rfl + +end with_lower_topology + +/-- +The lower topology is the topology generated by the complements of the closed intervals to infinity. +-/ +class lower_topology (α : Type*) [t : topological_space α] [preorder α] : Prop := +(topology_eq_lower_topology [] : t = generate_from {s | ∃ a, (Ici a)ᶜ = s}) + +instance [preorder α] : lower_topology (with_lower_topology α) := ⟨rfl⟩ + +namespace lower_topology + +/-- The complements of the upper closures of finite sets are a collection of lower sets +which form a basis for the lower topology. -/ +def lower_basis (α : Type*) [preorder α] := +{s : set α | ∃ t : set α, t.finite ∧ (upper_closure t : set α)ᶜ = s} + +section preorder +variables [preorder α] [topological_space α] [lower_topology α] {s : set α} + +/-- If `α` is equipped with the lower topology, then it is homeomorphic to `with_lower_topology α`. +-/ +def with_lower_topology_homeomorph : with_lower_topology α ≃ₜ α := +{ continuous_to_fun := by { convert continuous_id, apply topology_eq_lower_topology }, + continuous_inv_fun := by { convert ← continuous_id, apply topology_eq_lower_topology }, + ..with_lower_topology.of_lower } + +lemma is_open_iff_generate_Ici_compl : is_open s ↔ generate_open {t | ∃ a, (Ici a)ᶜ = t} s := +by rw topology_eq_lower_topology α; refl + +/-- Left-closed right-infinite intervals [a, ∞) are closed in the lower topology. -/ +lemma is_closed_Ici (a : α) : is_closed (Ici a) := +is_open_compl_iff.1 $ is_open_iff_generate_Ici_compl.2 $ generate_open.basic _ ⟨a, rfl⟩ + +/-- The upper closure of a finite set is closed in the lower topology. -/ +lemma is_closed_upper_closure (h : s.finite) : is_closed (upper_closure s : set α) := +begin + simp only [← upper_set.infi_Ici, upper_set.coe_infi], + exact is_closed_bUnion h (λ a h₁, is_closed_Ici a), +end + +/-- Every set open in the lower topology is a lower set. -/ +lemma is_lower_set_of_is_open (h : is_open s) : is_lower_set s := +begin + rw is_open_iff_generate_Ici_compl at h, + induction h, + case generate_open.basic : u h { obtain ⟨a, rfl⟩ := h, exact (is_upper_set_Ici a).compl }, + case univ : { exact is_lower_set_univ }, + case inter : u v hu1 hv1 hu2 hv2 { exact hu2.inter hv2 }, + case sUnion : _ _ ih { exact is_lower_set_sUnion ih }, +end + +lemma is_upper_set_of_is_closed (h : is_closed s) : is_upper_set s := +is_lower_set_compl.1 $ is_lower_set_of_is_open h.is_open_compl + +/-- +The closure of a singleton `{a}` in the lower topology is the left-closed right-infinite interval +[a, ∞). +-/ +@[simp] lemma closure_singleton (a : α) : closure {a} = Ici a := +subset_antisymm (closure_minimal (λ b h, h.ge) $ is_closed_Ici a) $ + (is_upper_set_of_is_closed is_closed_closure).Ici_subset $ subset_closure rfl + +protected lemma is_topological_basis : + is_topological_basis (lower_basis α) := +begin + convert is_topological_basis_of_subbasis (topology_eq_lower_topology α), + simp_rw [lower_basis, coe_upper_closure, compl_Union], + ext s, + split, + { rintro ⟨F, hF, rfl⟩, + refine ⟨(λ a, (Ici a)ᶜ) '' F, ⟨hF.image _, image_subset_iff.2 $ λ _ _, ⟨_, rfl⟩⟩, _⟩, + rw sInter_image }, + { rintro ⟨F, ⟨hF, hs⟩, rfl⟩, + haveI := hF.to_subtype, + rw [subset_def, subtype.forall'] at hs, + choose f hf using hs, + exact ⟨_, finite_range f, by simp_rw [bInter_range, hf, sInter_eq_Inter]⟩ } +end + +end preorder + +section partial_order +variables [partial_order α] [topological_space α] [lower_topology α] + +/-- +The lower topology on a partial order is T₀. +-/ +@[priority 90] -- see Note [lower instance priority] +instance : t0_space α := +(t0_space_iff_inseparable α).2 $ λ x y h, Ici_injective $ + by simpa only [inseparable_iff_closure_eq, closure_singleton] using h + +end partial_order +end lower_topology + +instance [preorder α] [topological_space α] [lower_topology α] [order_bot α] + [preorder β] [topological_space β] [lower_topology β] [order_bot β] : lower_topology (α × β) := +{ topology_eq_lower_topology := + begin + refine le_antisymm (le_generate_from _) _, + { rintro _ ⟨x, rfl⟩, + exact ((lower_topology.is_closed_Ici _).prod $ + lower_topology.is_closed_Ici _).is_open_compl }, + rw [(lower_topology.is_topological_basis.prod + lower_topology.is_topological_basis).eq_generate_from, + le_generate_from_iff_subset_is_open, image2_subset_iff], + rintro _ ⟨s, hs, rfl⟩ _ ⟨t, ht, rfl⟩, + dsimp, + simp_rw [coe_upper_closure, compl_Union, prod_eq, preimage_Inter, preimage_compl], + -- Note: `refine` doesn't work here because it tries using `prod.topological_space`. + apply (is_open_bInter hs $ λ a _, _).inter (is_open_bInter ht $ λ b _, _), + { exact generate_open.basic _ ⟨(a, ⊥), by simp [Ici_prod_eq, prod_univ]⟩ }, + { exact generate_open.basic _ ⟨(⊥, b), by simp [Ici_prod_eq, univ_prod]⟩ }, + all_goals { apply_instance }, + end } + +section complete_lattice +variables [complete_lattice α] [complete_lattice β] [topological_space α] [lower_topology α] + [topological_space β] [lower_topology β] + +lemma Inf_hom.continuous (f : Inf_hom α β) : continuous f := +begin + convert continuous_generated_from _, + { exact lower_topology.topology_eq_lower_topology β }, + rintro _ ⟨b, rfl⟩, + rw [preimage_compl, is_open_compl_iff], + convert lower_topology.is_closed_Ici (Inf $ f ⁻¹' Ici b), + refine subset_antisymm (λ a, Inf_le) (λ a ha, le_trans _ $ order_hom_class.mono f ha), + simp [map_Inf], +end + +@[priority 90] -- see Note [lower instance priority] +instance lower_topology.to_has_continuous_inf : has_continuous_inf α := +⟨(inf_Inf_hom : Inf_hom (α × α) α).continuous⟩ + +end complete_lattice From 2751ae214172a40dfdd28dcbe91767868e0d8b81 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Sat, 21 Jan 2023 12:48:09 +0000 Subject: [PATCH 0312/1291] refactor(set_theory/ordinal/arithmetic): fix `rank` universes (#18239) The rank of an object in universe `u` should also be an ordinal in universe `u`, rather than universe `max u v` for some other `v`, as `ordinal.lift` can achieve this effect. --- src/order/extension/well.lean | 2 +- src/set_theory/ordinal/arithmetic.lean | 19 ++++++++++--------- 2 files changed, 11 insertions(+), 10 deletions(-) diff --git a/src/order/extension/well.lean b/src/order/extension/well.lean index 76c290e56f5d0..b684cf4c87f44 100644 --- a/src/order/extension/well.lean +++ b/src/order/extension/well.lean @@ -44,7 +44,7 @@ arbitrary well-order to serve as a tiebreak between two elements of same rank. noncomputable def well_order_extension : linear_order α := let l : linear_order α := is_well_order.linear_order well_ordering_rel in by exactI @linear_order.lift' α (ordinal ×ₗ α) _ - (λ a : α, (well_founded.rank.{u u} hwf a, a)) (λ _ _, congr_arg prod.snd) + (λ a : α, (well_founded.rank.{u} hwf a, a)) (λ _ _, congr_arg prod.snd) instance well_order_extension.is_well_founded_lt : is_well_founded α hwf.well_order_extension.lt := ⟨inv_image.wf _ $ prod.lex_wf ordinal.well_founded_lt.wf well_ordering_rel.is_well_order.wf⟩ diff --git a/src/set_theory/ordinal/arithmetic.lean b/src/set_theory/ordinal/arithmetic.lean index 810d62bb70b87..6057463f2d35b 100644 --- a/src/set_theory/ordinal/arithmetic.lean +++ b/src/set_theory/ordinal/arithmetic.lean @@ -54,10 +54,10 @@ open function cardinal set equiv order open_locale classical cardinal ordinal universes u v w -variables {α : Type*} {β : Type*} {γ : Type*} - {r : α → α → Prop} {s : β → β → Prop} {t : γ → γ → Prop} namespace ordinal +variables {α : Type*} {β : Type*} {γ : Type*} + {r : α → α → Prop} {s : β → β → Prop} {t : γ → γ → Prop} /-! ### Further properties of addition on ordinals -/ @@ -2332,17 +2332,18 @@ meta def positivity_opow : expr → tactic strictness end tactic +variables {α : Type u} {r : α → α → Prop} {a b : α} + namespace acc -variables {a b : α} /-- The rank of an element `a` accessible under a relation `r` is defined inductively as the smallest ordinal greater than the ranks of all elements below it (i.e. elements `b` such that `r b a`). -/ -noncomputable def rank (h : acc r a) : ordinal := -acc.rec_on h $ λ a h ih, ordinal.sup $ λ b : {b // r b a}, order.succ $ ih b b.2 +noncomputable def rank (h : acc r a) : ordinal.{u} := +acc.rec_on h $ λ a h ih, ordinal.sup.{u u} $ λ b : {b // r b a}, order.succ $ ih b b.2 lemma rank_eq (h : acc r a) : - h.rank = ordinal.sup (λ b : {b // r b a}, order.succ (h.inv b.2).rank) := + h.rank = ordinal.sup.{u u} (λ b : {b // r b a}, order.succ (h.inv b.2).rank) := by { change (acc.intro a $ λ _, h.inv).rank = _, refl } /-- if `r a b` then the rank of `a` is less than the rank of `b`. -/ @@ -2352,15 +2353,15 @@ lemma rank_lt_of_rel (hb : acc r b) (h : r a b) : (hb.inv h).rank < hb.rank := end acc namespace well_founded -variables (hwf : well_founded r) {a b : α} +variables (hwf : well_founded r) include hwf /-- The rank of an element `a` under a well-founded relation `r` is defined inductively as the smallest ordinal greater than the ranks of all elements below it (i.e. elements `b` such that `r b a`). -/ -noncomputable def rank (a : α) : ordinal := (hwf.apply a).rank +noncomputable def rank (a : α) : ordinal.{u} := (hwf.apply a).rank -lemma rank_eq : hwf.rank a = ordinal.sup (λ b : {b // r b a}, order.succ $ hwf.rank b) := +lemma rank_eq : hwf.rank a = ordinal.sup.{u u} (λ b : {b // r b a}, order.succ $ hwf.rank b) := by { rw [rank, acc.rank_eq], refl } lemma rank_lt_of_rel (h : r a b) : hwf.rank a < hwf.rank b := acc.rank_lt_of_rel _ h From d23418e009d2bc37f9e1047ecaa229790b7356be Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 21 Jan 2023 14:48:25 +0000 Subject: [PATCH 0313/1291] feat(algebra/group/inj_surj): Missing transfer instances (#18247) Transfer `add_comm_monoid_with_one`/`add_comm_group_with_one` along injective/surjective functions. Also add a missing `reducible` tag. --- src/algebra/group/inj_surj.lean | 56 +++++++++++++++++++++++++++++++++ 1 file changed, 56 insertions(+) diff --git a/src/algebra/group/inj_surj.lean b/src/algebra/group/inj_surj.lean index d9c43e3513bdd..68f1acbe5576b 100644 --- a/src/algebra/group/inj_surj.lean +++ b/src/algebra/group/inj_surj.lean @@ -179,6 +179,17 @@ protected def comm_monoid [comm_monoid M₂] (f : M₁ → M₂) (hf : injective comm_monoid M₁ := { .. hf.comm_semigroup f mul, .. hf.monoid f one mul npow } +/-- A type endowed with `0`, `1` and `+` is an additive commutative monoid with one, if it admits an +injective map that preserves `0`, `1` and `+` to an additive commutative monoid with one. +See note [reducible non-instances]. -/ +@[reducible] +protected def add_comm_monoid_with_one {M₁} [has_zero M₁] [has_one M₁] [has_add M₁] [has_smul ℕ M₁] + [has_nat_cast M₁] [add_comm_monoid_with_one M₂] (f : M₁ → M₂) (hf : injective f) (zero : f 0 = 0) + (one : f 1 = 1) (add : ∀ x y, f (x + y) = f x + f y) (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) + (nat_cast : ∀ n : ℕ, f n = n) : + add_comm_monoid_with_one M₁ := +{ ..hf.add_monoid_with_one f zero one add nsmul nat_cast, ..hf.add_comm_monoid f zero add nsmul } + /-- A type endowed with `1` and `*` is a cancel commutative monoid, if it admits an injective map that preserves `1` and `*` to a cancel commutative monoid. See note [reducible non-instances]. -/ @@ -303,6 +314,21 @@ protected def comm_group [comm_group M₂] (f : M₁ → M₂) (hf : injective f comm_group M₁ := { .. hf.comm_monoid f one mul npow, .. hf.group f one mul inv div npow zpow } +/-- A type endowed with `0`, `1` and `+` is an additive commutative group with one, if it admits an +injective map that preserves `0`, `1` and `+` to an additive commutative group with one. +See note [reducible non-instances]. -/ +@[reducible] +protected def add_comm_group_with_one {M₁} [has_zero M₁] [has_one M₁] [has_add M₁] [has_smul ℕ M₁] + [has_neg M₁] [has_sub M₁] [has_smul ℤ M₁] [has_nat_cast M₁] [has_int_cast M₁] + [add_comm_group_with_one M₂] (f : M₁ → M₂) (hf : injective f) + (zero : f 0 = 0) (one : f 1 = 1) (add : ∀ x y, f (x + y) = f x + f y) + (neg : ∀ x, f (- x) = - f x) (sub : ∀ x y, f (x - y) = f x - f y) + (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (zsmul : ∀ x (n : ℤ), f (n • x) = n • f x) + (nat_cast : ∀ n : ℕ, f n = n) (int_cast : ∀ n : ℤ, f n = n) : + add_comm_group_with_one M₁ := +{ ..hf.add_group_with_one f zero one add neg sub nsmul zsmul nat_cast int_cast, + ..hf.add_comm_monoid f zero add nsmul } + end injective /-! @@ -395,6 +421,19 @@ protected def comm_monoid [comm_monoid M₁] (f : M₁ → M₂) (hf : surjectiv comm_monoid M₂ := { .. hf.comm_semigroup f mul, .. hf.monoid f one mul npow } +/-- A type endowed with `0`, `1` and `+` is an additive monoid with one, +if it admits a surjective map that preserves `0`, `1` and `*` from an additive monoid with one. +See note [reducible non-instances]. -/ +@[reducible] +protected def add_comm_monoid_with_one + {M₂} [has_zero M₂] [has_one M₂] [has_add M₂] [has_smul ℕ M₂] [has_nat_cast M₂] + [add_comm_monoid_with_one M₁] (f : M₁ → M₂) (hf : surjective f) + (zero : f 0 = 0) (one : f 1 = 1) (add : ∀ x y, f (x + y) = f x + f y) + (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) + (nat_cast : ∀ n : ℕ, f n = n) : + add_comm_monoid_with_one M₂ := +{ ..hf.add_monoid_with_one f zero one add nsmul nat_cast, ..hf.add_comm_monoid _ zero _ nsmul } + /-- A type has an involutive inversion if it admits a surjective map that preserves `⁻¹` to a type which has an involutive inversion. -/ @[reducible, to_additive "A type has an involutive negation if it admits a surjective map that @@ -446,6 +485,7 @@ protected def group [group M₁] (f : M₁ → M₂) (hf : surjective f) /-- A type endowed with `0`, `1`, `+` is an additive group with one, if it admits a surjective map that preserves `0`, `1`, and `+` to an additive group with one. See note [reducible non-instances]. -/ +@[reducible] protected def add_group_with_one {M₂} [has_zero M₂] [has_one M₂] [has_add M₂] [has_neg M₂] [has_sub M₂] [has_smul ℕ M₂] [has_smul ℤ M₂] [has_nat_cast M₂] [has_int_cast M₂] @@ -475,6 +515,22 @@ protected def comm_group [comm_group M₁] (f : M₁ → M₂) (hf : surjective comm_group M₂ := { .. hf.comm_monoid f one mul npow, .. hf.group f one mul inv div npow zpow } +/-- A type endowed with `0`, `1`, `+` is an additive commutative group with one, if it admits a +surjective map that preserves `0`, `1`, and `+` to an additive commutative group with one. +See note [reducible non-instances]. -/ +@[reducible] +protected def add_comm_group_with_one + {M₂} [has_zero M₂] [has_one M₂] [has_add M₂] [has_neg M₂] [has_sub M₂] + [has_smul ℕ M₂] [has_smul ℤ M₂] [has_nat_cast M₂] [has_int_cast M₂] + [add_comm_group_with_one M₁] (f : M₁ → M₂) (hf : surjective f) + (zero : f 0 = 0) (one : f 1 = 1) (add : ∀ x y, f (x + y) = f x + f y) + (neg : ∀ x, f (- x) = - f x) (sub : ∀ x y, f (x - y) = f x - f y) + (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (zsmul : ∀ x (n : ℤ), f (n • x) = n • f x) + (nat_cast : ∀ n : ℕ, f n = n) (int_cast : ∀ n : ℤ, f n = n) : + add_comm_group_with_one M₂ := +{ ..hf.add_group_with_one f zero one add neg sub nsmul zsmul nat_cast int_cast, + ..hf.add_comm_monoid _ zero add nsmul } + end surjective end function From e1503e02cbd97dfa07584c81b46cfc4f39be79d9 Mon Sep 17 00:00:00 2001 From: Kyle Miller Date: Sat, 21 Jan 2023 17:49:28 +0000 Subject: [PATCH 0314/1291] feat(data/sym/sym2): `set_like` instance (#17154) --- src/data/sym/sym2.lean | 42 +++++++++++++++++++++++++----------------- 1 file changed, 25 insertions(+), 17 deletions(-) diff --git a/src/data/sym/sym2.lean b/src/data/sym/sym2.lean index 153f56f273a38..f7287e12bd1cf 100644 --- a/src/data/sym/sym2.lean +++ b/src/data/sym/sym2.lean @@ -5,6 +5,7 @@ Authors: Kyle Miller -/ import data.finset.prod import data.sym.basic +import data.set_like.basic import tactic.linarith /-! @@ -199,24 +200,42 @@ end section membership -/-! ### Declarations about membership -/ +/-! ### Membership and set coercion -/ /-- This is a predicate that determines whether a given term is a member of a term of the symmetric square. From this point of view, the symmetric square is the subtype of cardinality-two multisets on `α`. -/ -def mem (x : α) (z : sym2 α) : Prop := +protected def mem (x : α) (z : sym2 α) : Prop := ∃ (y : α), z = ⟦(x, y)⟧ -instance : has_mem α (sym2 α) := ⟨mem⟩ +lemma mem_iff' {a b c : α} : sym2.mem a ⟦(b, c)⟧ ↔ a = b ∨ a = c := +{ mp := by { rintro ⟨_, h⟩, rw eq_iff at h, tidy }, + mpr := by { rintro (rfl|rfl), { exact ⟨_, rfl⟩ }, rw eq_swap, exact ⟨_, rfl⟩ } } + +instance : set_like (sym2 α) α := +{ coe := λ z, {x | z.mem x}, + coe_injective' := λ z z' h, begin + simp only [set.ext_iff, set.mem_set_of_eq] at h, + induction z using sym2.ind with x y, + induction z' using sym2.ind with x' y', + have hx := h x, have hy := h y, have hx' := h x', have hy' := h y', + simp only [mem_iff', eq_self_iff_true, or_true, iff_true, true_or, true_iff] at hx hy hx' hy', + cases hx; cases hy; cases hx'; cases hy'; subst_vars, + rw [sym2.eq_swap], + end } + +@[simp] lemma mem_iff_mem {x : α} {z : sym2 α} : sym2.mem x z ↔ x ∈ z := iff.rfl + +lemma mem_iff_exists {x : α} {z : sym2 α} : x ∈ z ↔ ∃ (y : α), z = ⟦(x, y)⟧ := iff.rfl + +@[ext] theorem ext {p q : sym2 α} (h : ∀ x, x ∈ p ↔ x ∈ q) : p = q := set_like.ext h lemma mem_mk_left (x y : α) : x ∈ ⟦(x, y)⟧ := ⟨y, rfl⟩ lemma mem_mk_right (x y : α) : y ∈ ⟦(x, y)⟧ := eq_swap.subst $ mem_mk_left y x -@[simp] lemma mem_iff {a b c : α} : a ∈ ⟦(b, c)⟧ ↔ a = b ∨ a = c := -{ mp := by { rintro ⟨_, h⟩, rw eq_iff at h, tidy }, - mpr := by { rintro ⟨_⟩; subst a, { apply mem_mk_left }, apply mem_mk_right } } +@[simp] lemma mem_iff {a b c : α} : a ∈ ⟦(b, c)⟧ ↔ a = b ∨ a = c := mem_iff' lemma out_fst_mem (e : sym2 α) : e.out.1 ∈ e := ⟨e.out.2, by rw [prod.mk.eta, e.out_eq]⟩ lemma out_snd_mem (e : sym2 α) : e.out.2 ∈ e := ⟨e.out.1, by rw [eq_swap, prod.mk.eta, e.out_eq]⟩ @@ -258,17 +277,6 @@ lemma eq_of_ne_mem {x y : α} {z z' : sym2 α} (h : x ≠ y) (h1 : x ∈ z) (h2 : y ∈ z) (h3 : x ∈ z') (h4 : y ∈ z') : z = z' := ((mem_and_mem_iff h).mp ⟨h1, h2⟩).trans ((mem_and_mem_iff h).mp ⟨h3, h4⟩).symm -@[ext] -protected lemma ext (z z' : sym2 α) (h : ∀ x, x ∈ z ↔ x ∈ z') : z = z' := -begin - induction z using sym2.ind with x y, - induction z' using sym2.ind with x' y', - have hx := h x, have hy := h y, have hx' := h x', have hy' := h y', - simp only [mem_iff, eq_self_iff_true, or_true, iff_true, true_or, true_iff] at hx hy hx' hy', - cases hx; cases hy; cases hx'; cases hy'; subst_vars, - simp only [sym2.eq_swap], -end - instance mem.decidable [decidable_eq α] (x : α) (z : sym2 α) : decidable (x ∈ z) := quotient.rec_on_subsingleton z (λ ⟨y₁, y₂⟩, decidable_of_iff' _ mem_iff) From d6fad0e5bf2d6f48da9175d25c3dc5706b3834ce Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9mi=20Bottinelli?= Date: Sat, 21 Jan 2023 19:28:48 +0000 Subject: [PATCH 0315/1291] feat(analysis/bounded_variation): define `variation_on_from_to` (#18040) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit … and refactor `has_locally_bounded_variation_on.exists_monotone_on_sub_monotone_on` using it. Co-authored-by: Junyan Xu Co-authored-by: Rémi Bottinelli --- src/analysis/bounded_variation.lean | 159 ++++++++++++++++++++-------- 1 file changed, 113 insertions(+), 46 deletions(-) diff --git a/src/analysis/bounded_variation.lean b/src/analysis/bounded_variation.lean index bb435c8f434af..aa78a0f8eda05 100644 --- a/src/analysis/bounded_variation.lean +++ b/src/analysis/bounded_variation.lean @@ -6,6 +6,7 @@ Authors: Sébastien Gouëzel import measure_theory.measure.lebesgue import analysis.calculus.monotone import data.set.function +import algebra.group.basic /-! # Functions of bounded variation @@ -21,6 +22,7 @@ almost everywhere. * `has_bounded_variation_on f s` registers that the variation of `f` on `s` is finite. * `has_locally_bounded_variation f s` registers that `f` has finite variation on any compact subinterval of `s`. +* `variation_on_from_to f s a b` is the signed variation of `f` on `s ∩ Icc a b`, converted to `ℝ`. * `evariation_on.Icc_add_Icc` states that the variation of `f` on `[a, c]` is the sum of its variations on `[a, b]` and `[b, c]`. @@ -686,60 +688,126 @@ lemma monotone_on.has_locally_bounded_variation_on {f : α → ℝ} {s : set α} has_locally_bounded_variation_on f s := λ a b as bs, ((hf.evariation_on_le as bs).trans_lt ennreal.of_real_lt_top).ne +/-- +The **signed** variation of `f` on the interval `Icc a b` intersected with the set `s`, +squashed to a real (therefore only really meaningful if the variation is finite) +-/ +noncomputable def variation_on_from_to (f : α → E) (s : set α) (a b : α) : ℝ := +if a ≤ b then (evariation_on f (s ∩ Icc a b)).to_real else + - (evariation_on f (s ∩ Icc b a)).to_real + +namespace variation_on_from_to + +variables (f : α → E) (s : set α) + +@[protected] +lemma self (a : α) : variation_on_from_to f s a a = 0 := +begin + dsimp only [variation_on_from_to], + rw [if_pos le_rfl, Icc_self, evariation_on.subsingleton, ennreal.zero_to_real], + exact λ x hx y hy, hx.2.trans hy.2.symm, +end + +@[protected] +lemma nonneg_of_le {a b : α} (h : a ≤ b) : 0 ≤ variation_on_from_to f s a b := +by simp only [variation_on_from_to, if_pos h, ennreal.to_real_nonneg] + +@[protected] +lemma eq_neg_swap (a b : α) : + variation_on_from_to f s a b = - variation_on_from_to f s b a := +begin + rcases lt_trichotomy a b with ab|rfl|ba, + { simp only [variation_on_from_to, if_pos ab.le, if_neg ab.not_le, neg_neg], }, + { simp only [self, neg_zero], }, + { simp only [variation_on_from_to, if_pos ba.le, if_neg ba.not_le, neg_neg], }, +end + +@[protected] +lemma nonpos_of_ge {a b : α} (h : b ≤ a) : variation_on_from_to f s a b ≤ 0 := +begin + rw eq_neg_swap, + exact neg_nonpos_of_nonneg (nonneg_of_le f s h), +end + +@[protected] +lemma eq_of_le {a b : α} (h : a ≤ b) : + variation_on_from_to f s a b = (evariation_on f (s ∩ Icc a b)).to_real := if_pos h + +@[protected] +lemma eq_of_ge {a b : α} (h : b ≤ a) : + variation_on_from_to f s a b = - (evariation_on f (s ∩ Icc b a)).to_real := +by rw [eq_neg_swap, neg_inj, eq_of_le f s h] + +@[protected] +lemma add {f : α → E} {s : set α} (hf : has_locally_bounded_variation_on f s) + {a b c : α} (ha : a ∈ s) (hb : b ∈ s) (hc : c ∈ s) : + variation_on_from_to f s a b + variation_on_from_to f s b c = variation_on_from_to f s a c := +begin + symmetry, + refine additive_of_is_total (≤) (variation_on_from_to f s) (∈s) _ _ ha hb hc, + { rintro x y xs ys, + simp only [eq_neg_swap f s y x, subtype.coe_mk, add_right_neg, forall_true_left], }, + { rintro x y z xy yz xs ys zs, + rw [eq_of_le f s xy, eq_of_le f s yz, eq_of_le f s (xy.trans yz), + ←ennreal.to_real_add (hf x y xs ys) (hf y z ys zs), + evariation_on.Icc_add_Icc f xy yz ys], }, +end + +variables {f} {s} + +@[protected] +lemma monotone_on (hf : has_locally_bounded_variation_on f s) + {a : α} (as : a ∈ s) : monotone_on (variation_on_from_to f s a) s := +begin + rintro b bs c cs bc, + rw ←add hf as bs cs, + exact le_add_of_nonneg_right (nonneg_of_le f s bc), +end + +@[protected] +lemma antitone_on (hf : has_locally_bounded_variation_on f s) + {b : α} (bs : b ∈ s) : antitone_on (λ a, variation_on_from_to f s a b) s := +begin + rintro a as c cs ac, + dsimp only, + rw ←add hf as cs bs, + exact le_add_of_nonneg_left (nonneg_of_le f s ac), +end + +@[protected] +lemma sub_self_monotone_on {f : α → ℝ} {s : set α} + (hf : has_locally_bounded_variation_on f s) {a : α} (as : a ∈ s) : + monotone_on (variation_on_from_to f s a - f) s := +begin + rintro b bs c cs bc, + rw [pi.sub_apply, pi.sub_apply, le_sub_iff_add_le, add_comm_sub, ← le_sub_iff_add_le'], + calc f c - f b + ≤ |f c - f b| : le_abs_self _ + ... = dist (f b) (f c) : by rw [dist_comm, real.dist_eq] + ... ≤ variation_on_from_to f s b c : by + { rw [eq_of_le f s bc, dist_edist], + apply ennreal.to_real_mono (hf b c bs cs), + apply evariation_on.edist_le f, + exacts [⟨bs, le_rfl, bc⟩, ⟨cs, bc, le_rfl⟩] } + ... = variation_on_from_to f s a c - variation_on_from_to f s a b : + by rw [←add hf as bs cs, add_sub_cancel'] +end + +end variation_on_from_to + /-- If a real valued function has bounded variation on a set, then it is a difference of monotone functions there. -/ lemma has_locally_bounded_variation_on.exists_monotone_on_sub_monotone_on {f : α → ℝ} {s : set α} (h : has_locally_bounded_variation_on f s) : ∃ (p q : α → ℝ), monotone_on p s ∧ monotone_on q s ∧ f = p - q := begin - rcases eq_empty_or_nonempty s with rfl|hs, + rcases eq_empty_or_nonempty s with rfl|⟨c, cs⟩, { exact ⟨f, 0, subsingleton_empty.monotone_on _, subsingleton_empty.monotone_on _, - by simp only [tsub_zero]⟩ }, - rcases hs with ⟨c, cs⟩, - let p := λ x, if c ≤ x then (evariation_on f (s ∩ Icc c x)).to_real - else -(evariation_on f (s ∩ Icc x c)).to_real, - have hp : monotone_on p s, - { assume x xs y ys hxy, - dsimp only [p], - split_ifs with hcx hcy hcy, - { have : evariation_on f (s ∩ Icc c x) + evariation_on f (s ∩ Icc x y) - = evariation_on f (s ∩ Icc c y), from evariation_on.Icc_add_Icc f hcx hxy xs, - rw [← this, ennreal.to_real_add (h c x cs xs) (h x y xs ys)], - exact le_add_of_le_of_nonneg le_rfl ennreal.to_real_nonneg }, - { exact (lt_irrefl _ ((not_le.1 hcy).trans_le (hcx.trans hxy))).elim }, - { exact (neg_nonpos.2 ennreal.to_real_nonneg).trans ennreal.to_real_nonneg }, - { simp only [neg_le_neg_iff], - have : evariation_on f (s ∩ Icc x y) + evariation_on f (s ∩ Icc y c) - = evariation_on f (s ∩ Icc x c), from evariation_on.Icc_add_Icc f hxy (not_le.1 hcy).le ys, - rw [← this, ennreal.to_real_add (h x y xs ys) (h y c ys cs), add_comm], - exact le_add_of_le_of_nonneg le_rfl ennreal.to_real_nonneg } }, - have hq : monotone_on (λ x, p x - f x) s, - { assume x xs y ys hxy, - dsimp only [p], - split_ifs with hcx hcy hcy, - { have : evariation_on f (s ∩ Icc c x) + evariation_on f (s ∩ Icc x y) - = evariation_on f (s ∩ Icc c y), from evariation_on.Icc_add_Icc f hcx hxy xs, - rw [← this, ennreal.to_real_add (h c x cs xs) (h x y xs ys)], - suffices : f y - f x ≤ (evariation_on f (s ∩ Icc x y)).to_real, by linarith, - exact (h x y xs ys).sub_le ⟨ys, hxy, le_rfl⟩ ⟨xs, le_rfl, hxy⟩ }, - { exact (lt_irrefl _ ((not_le.1 hcy).trans_le (hcx.trans hxy))).elim }, - { suffices : f y - f x ≤ (evariation_on f (s ∩ Icc x c)).to_real - + (evariation_on f (s ∩ Icc c y)).to_real, by linarith, - rw [← ennreal.to_real_add (h x c xs cs) (h c y cs ys), - evariation_on.Icc_add_Icc f (not_le.1 hcx).le hcy cs], - exact (h x y xs ys).sub_le ⟨ys, hxy, le_rfl⟩ ⟨xs, le_rfl, hxy⟩ }, - { have : evariation_on f (s ∩ Icc x y) + evariation_on f (s ∩ Icc y c) - = evariation_on f (s ∩ Icc x c), from evariation_on.Icc_add_Icc f hxy (not_le.1 hcy).le ys, - rw [← this, ennreal.to_real_add (h x y xs ys) (h y c ys cs)], - suffices : f y - f x ≤ (evariation_on f (s ∩ Icc x y)).to_real, by linarith, - exact (h x y xs ys).sub_le ⟨ys, hxy, le_rfl⟩ ⟨xs, le_rfl, hxy⟩ } }, - refine ⟨p, λ x, p x - f x, hp, hq, _⟩, - ext x, - dsimp, - abel, + (sub_zero f).symm⟩ }, + { exact ⟨_, _, variation_on_from_to.monotone_on h cs, + variation_on_from_to.sub_self_monotone_on h cs, (sub_sub_cancel _ _).symm⟩ }, end - /-! ## Lipschitz functions and bounded variation -/ lemma lipschitz_on_with.comp_evariation_on_le {f : E → F} {C : ℝ≥0} {t : set E} @@ -885,4 +953,3 @@ lemma lipschitz_with.ae_differentiable_at {C : ℝ≥0} {f : ℝ → V} (h : lipschitz_with C f) : ∀ᵐ x, differentiable_at ℝ f x := (h.has_locally_bounded_variation_on univ).ae_differentiable_at - From f2edd790f6c6e1d660515f76768f63cb717434d7 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sun, 22 Jan 2023 01:07:31 +0000 Subject: [PATCH 0316/1291] fix(linear_algebra/basis): add missing decidable arguments in lemmas (#18252) The resulting lemmas are syntactically more general. To ensure that this catches all of them, this also removes `open_locale classical` from the file, replacing it with manual construction of classical decidability instances within definitions, and `classical` within proofs. --- src/analysis/normed_space/pi_Lp.lean | 3 +-- src/linear_algebra/basis.lean | 29 ++++++++++++++++++---------- src/linear_algebra/std_basis.lean | 4 +--- 3 files changed, 21 insertions(+), 15 deletions(-) diff --git a/src/analysis/normed_space/pi_Lp.lean b/src/analysis/normed_space/pi_Lp.lean index 12b33f1a0c6db..acb6eda1527de 100644 --- a/src/analysis/normed_space/pi_Lp.lean +++ b/src/analysis/normed_space/pi_Lp.lean @@ -698,8 +698,7 @@ basis.of_equiv_fun (pi_Lp.linear_equiv p 𝕜 (λ _ : ι, 𝕜)) @[simp] lemma basis_fun_apply [decidable_eq ι] (i) : basis_fun p 𝕜 ι i = (pi_Lp.equiv p _).symm (pi.single i 1) := -by { simp_rw [basis_fun, basis.coe_of_equiv_fun, pi_Lp.linear_equiv_symm_apply, pi.single], - congr /- Get rid of a `decidable_eq` mismatch. -/ } +by simp_rw [basis_fun, basis.coe_of_equiv_fun, pi_Lp.linear_equiv_symm_apply, pi.single] @[simp] lemma basis_fun_repr (x : pi_Lp p (λ i : ι, 𝕜)) (i : ι) : (basis_fun p 𝕜 ι).repr x i = x i := diff --git a/src/linear_algebra/basis.lean b/src/linear_algebra/basis.lean index 4baa8870d7ddd..215594ec368a7 100644 --- a/src/linear_algebra/basis.lean +++ b/src/linear_algebra/basis.lean @@ -66,7 +66,7 @@ noncomputable theory universe u open function set submodule -open_locale classical big_operators +open_locale big_operators variables {ι : Type*} {ι' : Type*} {R : Type*} {R₂ : Type*} {K : Type*} variables {M : Type*} {M' M'' : Type*} {V : Type u} {V' : Type*} @@ -367,6 +367,7 @@ by rw [coe_reindex, range_reindex'] /-- `b.reindex_range` is a basis indexed by `range b`, the basis vectors themselves. -/ def reindex_range : basis (range b) R M := +by haveI := classical.dec (nontrivial R); exact if h : nontrivial R then by letI := h; exact b.reindex (equiv.of_injective b (basis.injective b)) else @@ -419,7 +420,7 @@ b.reindex_range_repr' _ rfl section fintype -variables [fintype ι] +variables [fintype ι] [decidable_eq M] /-- `b.reindex_finset_range` is a basis indexed by `finset.univ.image b`, the finite set of basis vectors themselves. -/ @@ -750,11 +751,13 @@ linear_equiv.trans b.repr ..finsupp.equiv_fun_on_finite } : (ι →₀ R) ≃ₗ[R] (ι → R)) /-- A module over a finite ring that admits a finite basis is finite. -/ -def module.fintype_of_fintype [fintype R] : fintype M := -fintype.of_equiv _ b.equiv_fun.to_equiv.symm +def module.fintype_of_fintype (b : basis ι R M) [fintype R] : fintype M := +by haveI := classical.dec_eq ι; exact + fintype.of_equiv _ b.equiv_fun.to_equiv.symm -theorem module.card_fintype [fintype R] [fintype M] : +theorem module.card_fintype (b : basis ι R M) [fintype R] [fintype M] : card M = (card R) ^ (card ι) := +by classical; exact calc card M = card (ι → R) : card_congr b.equiv_fun.to_equiv ... = card R ^ card ι : card_fun @@ -781,7 +784,8 @@ lemma basis.sum_repr (u : M) : ∑ i, b.repr u i • b i = u := b.sum_equiv_fun u @[simp] -lemma basis.equiv_fun_self (i j : ι) : b.equiv_fun (b i) j = if i = j then 1 else 0 := +lemma basis.equiv_fun_self [decidable_eq ι] (i j : ι) : + b.equiv_fun (b i) j = if i = j then 1 else 0 := by { rw [b.equiv_fun_apply, b.repr_self_apply] } lemma basis.repr_sum_self (c : ι → R) : ⇑(b.repr (∑ i, c i • b i)) = c := @@ -802,7 +806,7 @@ basis.of_repr $ e.trans $ linear_equiv.symm $ finsupp.linear_equiv_fun_on_finite @[simp] lemma basis.of_equiv_fun_repr_apply (e : M ≃ₗ[R] (ι → R)) (x : M) (i : ι) : (basis.of_equiv_fun e).repr x i = e x i := rfl -@[simp] lemma basis.coe_of_equiv_fun (e : M ≃ₗ[R] (ι → R)) : +@[simp] lemma basis.coe_of_equiv_fun [decidable_eq ι] (e : M ≃ₗ[R] (ι → R)) : (basis.of_equiv_fun e : ι → M) = λ i, e.symm (function.update 0 i 1) := funext $ λ i, e.injective $ funext $ λ j, by simp [basis.of_equiv_fun, ←finsupp.single_eq_pi_single, finsupp.single_eq_update] @@ -810,6 +814,7 @@ funext $ λ i, e.injective $ funext $ λ j, @[simp] lemma basis.of_equiv_fun_equiv_fun (v : basis ι R M) : basis.of_equiv_fun v.equiv_fun = v := begin + classical, ext j, simp only [basis.equiv_fun_symm_apply, basis.coe_of_equiv_fun], simp_rw [function.update_apply, ite_smul], @@ -971,7 +976,7 @@ by simp [hli.repr_eq_single j, h] /-- Given a basis, the `i`th element of the dual basis evaluates to the Kronecker delta on the `j`th element of the basis. -/ -lemma mk_coord_apply {i j : ι} : +lemma mk_coord_apply [decidable_eq ι] {i j : ι} : (basis.mk hli hsp).coord i (v j) = if j = i then 1 else 0 := begin cases eq_or_ne j i, @@ -1058,6 +1063,7 @@ mk_apply @[simp] lemma coord_units_smul (e : basis ι R₂ M) (w : ι → R₂ˣ) (i : ι) : (e.units_smul w).coord i = (w i)⁻¹ • e.coord i := begin + classical, apply e.ext, intros j, transitivity ((e.units_smul w).coord i) ((w j)⁻¹ • (e.units_smul w) j), @@ -1227,7 +1233,9 @@ let s := set.range v, b := hs.to_subtype_range.extend (subset_univ (set.range v)) in (basis.extend hs.to_subtype_range).reindex $ equiv.symm $ calc ι ⊕ (b \ s : set V) ≃ s ⊕ (b \ s : set V) : equiv.sum_congr e (equiv.refl _) - ... ≃ b : equiv.set.sum_diff_subset (hs.to_subtype_range.subset_extend _) + ... ≃ b : + by haveI := classical.dec_pred (∈ s); exact + equiv.set.sum_diff_subset (hs.to_subtype_range.subset_extend _) lemma subset_extend {s : set V} (hs : linear_independent K (coe : s → V)) : s ⊆ hs.extend (set.subset_univ _) := @@ -1275,7 +1283,8 @@ variables (K V) theorem vector_space.card_fintype [fintype K] [fintype V] : ∃ n : ℕ, card V = (card K) ^ n := -⟨card (basis.of_vector_space_index K V), module.card_fintype (basis.of_vector_space K V)⟩ +by classical; exact + ⟨card (basis.of_vector_space_index K V), module.card_fintype (basis.of_vector_space K V)⟩ section atoms_of_submodule_lattice diff --git a/src/linear_algebra/std_basis.lean b/src/linear_algebra/std_basis.lean index 1b47691c937d7..c198c0bcd33be 100644 --- a/src/linear_algebra/std_basis.lean +++ b/src/linear_algebra/std_basis.lean @@ -34,7 +34,6 @@ this is a basis over `fin 3 → R`. open function submodule open_locale big_operators -open_locale big_operators namespace linear_map @@ -247,8 +246,7 @@ basis.of_equiv_fun (linear_equiv.refl _ _) @[simp] lemma basis_fun_apply [decidable_eq η] (i) : basis_fun R η i = std_basis R (λ (i : η), R) i 1 := by { simp only [basis_fun, basis.coe_of_equiv_fun, linear_equiv.refl_symm, - linear_equiv.refl_apply, std_basis_apply], - congr /- Get rid of a `decidable_eq` mismatch. -/ } + linear_equiv.refl_apply, std_basis_apply] } @[simp] lemma basis_fun_repr (x : η → R) (i : η) : (pi.basis_fun R η).repr x i = x i := From 02315164da5cb771133aa7b1c297ab279d7d3464 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sun, 22 Jan 2023 07:21:32 +0000 Subject: [PATCH 0317/1291] refactor(order/filter/basic): drop a lemma (#18254) - Drop `filter.subtype_coe_map_comap_prod`. This lemma was used once in `mathlib` and Lean 4 has no `list_pair` instance that is used in its statement. - Add `uniformity_set_coe`. For some reason, Lean fails to rewrite on `uniformity_subtype` in this context. - Fix&golf a proof. --- src/order/filter/basic.lean | 5 ----- src/topology/uniform_space/basic.lean | 12 +++++++----- 2 files changed, 7 insertions(+), 10 deletions(-) diff --git a/src/order/filter/basic.lean b/src/order/filter/basic.lean index e4e33dbfccce4..ca2258da2cbd6 100644 --- a/src/order/filter/basic.lean +++ b/src/order/filter/basic.lean @@ -1880,11 +1880,6 @@ lemma subtype_coe_map_comap (s : set α) (f : filter α) : map (coe : s → α) (comap (coe : s → α) f) = f ⊓ 𝓟 s := by rw [map_comap, subtype.range_coe] -lemma subtype_coe_map_comap_prod (s : set α) (f : filter (α × α)) : - map (coe : s × s → α × α) (comap (coe : s × s → α × α) f) = f ⊓ 𝓟 (s ×ˢ s) := -have (coe : s × s → α × α) = (λ x, (x.1, x.2)), by ext ⟨x, y⟩; refl, -by simp [this, map_comap, ← prod_range_range_eq] - lemma image_mem_of_mem_comap {f : filter α} {c : β → α} (h : range c ∈ f) {W : set β} (W_in : W ∈ comap c f) : c '' W ∈ f := begin diff --git a/src/topology/uniform_space/basic.lean b/src/topology/uniform_space/basic.lean index 972b23fb70b02..c0f70498831b3 100644 --- a/src/topology/uniform_space/basic.lean +++ b/src/topology/uniform_space/basic.lean @@ -1340,6 +1340,10 @@ lemma uniformity_subtype {p : α → Prop} [t : uniform_space α] : 𝓤 (subtype p) = comap (λq:subtype p × subtype p, (q.1.1, q.2.1)) (𝓤 α) := rfl +lemma uniformity_set_coe {s : set α} [t : uniform_space α] : + 𝓤 s = comap (prod.map (coe : s → α) (coe : s → α)) (𝓤 α) := +rfl + lemma uniform_continuous_subtype_val {p : α → Prop} [uniform_space α] : uniform_continuous (subtype.val : {a : α // p a} → α) := uniform_continuous_comap @@ -1358,11 +1362,9 @@ lemma uniform_continuous_on_iff_restrict [uniform_space α] [uniform_space β] { uniform_continuous_on f s ↔ uniform_continuous (s.restrict f) := begin unfold uniform_continuous_on set.restrict uniform_continuous tendsto, - rw [show (λ x : s × s, (f x.1, f x.2)) = prod.map f f ∘ coe, by ext x; cases x; refl, - uniformity_comap rfl, - show prod.map subtype.val subtype.val = (coe : s × s → α × α), by ext x; cases x; refl], - conv in (map _ (comap _ _)) { rw ← filter.map_map }, - rw subtype_coe_map_comap_prod, refl, + conv_rhs { rw [show (λ x : s × s, (f x.1, f x.2)) = prod.map f f ∘ prod.map coe coe, from rfl, + uniformity_set_coe, ← map_map, map_comap, range_prod_map, subtype.range_coe] }, + refl end lemma tendsto_of_uniform_continuous_subtype From 831c494092374cfe9f50591ed0ac81a25efc5b86 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sun, 22 Jan 2023 09:26:01 +0000 Subject: [PATCH 0318/1291] refactor(group_theory/monoid_localization): flip the direction of multiplication (#15584) In a future PR, this will allow `localization` and `module_localization` to be defeq. Mathematically this makes no difference. For now, only the primed `localization.r'` is defeq to `localized_module.r`. Co-authored-by: Jujian Zhang --- src/algebra/module/localized_module.lean | 36 +++++---- src/algebraic_geometry/Spec.lean | 2 +- .../morphisms/quasi_separated.lean | 10 +-- .../projective_spectrum/scheme.lean | 23 +++--- src/algebraic_geometry/structure_sheaf.lean | 15 ++-- src/data/polynomial/laurent.lean | 2 +- src/field_theory/ratfunc.lean | 50 +++++------- src/group_theory/monoid_localization.lean | 78 +++++++++++-------- src/number_theory/cyclotomic/basic.lean | 2 +- src/number_theory/padics/padic_integers.lean | 2 +- src/ring_theory/artinian.lean | 4 +- .../homogeneous_localization.lean | 8 +- src/ring_theory/laurent_series.lean | 9 +-- src/ring_theory/local_properties.lean | 23 +++--- src/ring_theory/localization/at_prime.lean | 10 +-- src/ring_theory/localization/away.lean | 4 +- src/ring_theory/localization/basic.lean | 41 ++++++---- .../localization/fraction_ring.lean | 4 +- src/ring_theory/localization/ideal.lean | 7 +- src/ring_theory/localization/integral.lean | 4 +- .../localization_localization.lean | 27 +++---- src/ring_theory/localization/num_denom.lean | 2 +- src/ring_theory/ore_localization/basic.lean | 9 ++- src/ring_theory/ring_hom/surjective.lean | 8 +- .../valuation/valuation_subring.lean | 2 +- src/set_theory/surreal/dyadic.lean | 4 +- 26 files changed, 207 insertions(+), 179 deletions(-) diff --git a/src/algebra/module/localized_module.lean b/src/algebra/module/localized_module.lean index 6c33c39d43e93..8941bab7c10a4 100644 --- a/src/algebra/module/localized_module.lean +++ b/src/algebra/module/localized_module.lean @@ -46,26 +46,30 @@ variables (M : Type v) [add_comm_monoid M] [module R M] /--The equivalence relation on `M × S` where `(m1, s1) ≈ (m2, s2)` if and only if for some (u : S), u * (s2 • m1 - s1 • m2) = 0-/ -def r : (M × S) → (M × S) → Prop -| ⟨m1, s1⟩ ⟨m2, s2⟩ := ∃ (u : S), u • s1 • m2 = u • s2 • m1 +def r (a b : M × S) : Prop := +∃ (u : S), u • b.2 • a.1 = u • a.2 • b.1 lemma r.is_equiv : is_equiv _ (r S M) := { refl := λ ⟨m, s⟩, ⟨1, by rw [one_smul]⟩, trans := λ ⟨m1, s1⟩ ⟨m2, s2⟩ ⟨m3, s3⟩ ⟨u1, hu1⟩ ⟨u2, hu2⟩, begin use u1 * u2 * s2, -- Put everything in the same shape, sorting the terms using `simp` - have hu1' := congr_arg ((•) (u2 * s3)) hu1, - have hu2' := congr_arg ((•) (u1 * s1)) hu2, + have hu1' := congr_arg ((•) (u2 * s3)) hu1.symm, + have hu2' := congr_arg ((•) (u1 * s1)) hu2.symm, simp only [← mul_smul, smul_assoc, mul_assoc, mul_comm, mul_left_comm] at ⊢ hu1' hu2', rw [hu2', hu1'] end, symm := λ ⟨m1, s1⟩ ⟨m2, s2⟩ ⟨u, hu⟩, ⟨u, hu.symm⟩ } - instance r.setoid : setoid (M × S) := { r := r S M, iseqv := ⟨(r.is_equiv S M).refl, (r.is_equiv S M).symm, (r.is_equiv S M).trans⟩ } +-- TODO: change `localization` to use `r'` instead of `r` so that the two types are also defeq, +-- `localization S = localized_module S R`. +example {R} [comm_semiring R] (S : submonoid R) : ⇑(localization.r' S) = localized_module.r S R := +rfl + /-- If `S` is a multiplicative subset of a ring `R` and `M` an `R`-module, then we can localize `M` by `S`. @@ -80,7 +84,7 @@ variables {M S} def mk (m : M) (s : S) : localized_module S M := quotient.mk ⟨m, s⟩ -lemma mk_eq {m m' : M} {s s' : S} : mk m s = mk m' s' ↔ ∃ (u : S), u • s • m' = u • s' • m := +lemma mk_eq {m m' : M} {s s' : S} : mk m s = mk m' s' ↔ ∃ (u : S), u • s' • m = u • s • m' := quotient.eq @[elab_as_eliminator] @@ -149,7 +153,7 @@ begin rw [one_smul, one_smul], congr' 1, { rw [mul_assoc] }, - { rw [mul_comm, add_assoc, mul_smul, mul_smul, ←mul_smul sx sz, mul_comm, mul_smul], }, + { rw [eq_comm, mul_comm, add_assoc, mul_smul, mul_smul, ←mul_smul sx sz, mul_comm, mul_smul], }, end private lemma add_comm' (x y : localized_module S M) : @@ -207,9 +211,10 @@ instance {A : Type*} [semiring A] [algebra R A] {S : submonoid R} : rintros ⟨a₁, s₁⟩ ⟨a₂, s₂⟩ ⟨b₁, t₁⟩ ⟨b₂, t₂⟩ ⟨u₁, e₁⟩ ⟨u₂, e₂⟩, rw mk_eq, use u₁ * u₂, - dsimp only, + dsimp only at ⊢ e₁ e₂, + rw eq_comm, transitivity (u₁ • t₁ • a₁) • u₂ • t₂ • a₂, - rw [← e₁, ← e₂], swap, rw eq_comm, + rw [e₁, e₂], swap, rw eq_comm, all_goals { rw [smul_smul, mul_mul_mul_comm, ← smul_eq_mul, ← smul_eq_mul A, smul_smul_smul_comm, mul_smul, mul_smul] } end), @@ -655,8 +660,9 @@ instance localized_module_is_localized_module : localized_module.mk_cancel t ], end, eq_iff_exists := λ m1 m2, - { mp := λ eq1, by simpa only [one_smul] using localized_module.mk_eq.mp eq1, - mpr := λ ⟨c, eq1⟩, localized_module.mk_eq.mpr ⟨c, by simpa only [one_smul] using eq1⟩ } } + { mp := λ eq1, by simpa only [eq_comm, one_smul] using localized_module.mk_eq.mp eq1, + mpr := λ ⟨c, eq1⟩, + localized_module.mk_eq.mpr ⟨c, by simpa only [eq_comm, one_smul] using eq1⟩ } } namespace is_localized_module @@ -674,7 +680,7 @@ begin generalize_proofs h1 h2, erw [module.End_algebra_map_is_unit_inv_apply_eq_iff, ←h2.unit⁻¹.1.map_smul, module.End_algebra_map_is_unit_inv_apply_eq_iff', ←linear_map.map_smul, ←linear_map.map_smul], - exact ((is_localized_module.eq_iff_exists S f).mpr ⟨c, eq1⟩).symm, + exact (is_localized_module.eq_iff_exists S f).mpr ⟨c, eq1⟩, end @[simp] lemma from_localized_module'_mk (m : M) (s : S) : @@ -894,7 +900,11 @@ by { rw [mk'_add_mk', ← smul_add, mk'_cancel_left] } lemma mk'_eq_mk'_iff (m₁ m₂ : M) (s₁ s₂ : S) : mk' f m₁ s₁ = mk' f m₂ s₂ ↔ ∃ s : S, s • s₁ • m₂ = s • s₂ • m₁ := -by { delta mk', rw [(from_localized_module.inj S f).eq_iff, localized_module.mk_eq] } +begin + delta mk', + rw [(from_localized_module.inj S f).eq_iff, localized_module.mk_eq], + simp_rw eq_comm +end lemma mk'_neg {M M' : Type*} [add_comm_group M] [add_comm_group M'] [module R M] [module R M'] (f : M →ₗ[R] M') [is_localized_module S f] (m : M) (s : S) : diff --git a/src/algebraic_geometry/Spec.lean b/src/algebraic_geometry/Spec.lean index fa176f103806e..95d525a67666d 100644 --- a/src/algebraic_geometry/Spec.lean +++ b/src/algebraic_geometry/Spec.lean @@ -353,7 +353,7 @@ begin _ _ _ (structure_sheaf.is_localization.to_basic_open S $ algebra_map R S r) x).trans this, obtain ⟨⟨_, n, rfl⟩, e⟩ := (is_localization.mk'_eq_zero_iff _ _).mp this, refine ⟨⟨r, hpr⟩ ^ n, _⟩, - rw [submonoid.smul_def, algebra.smul_def, submonoid.coe_pow, subtype.coe_mk, mul_comm, map_pow], + rw [submonoid.smul_def, algebra.smul_def, submonoid.coe_pow, subtype.coe_mk, map_pow], exact e }, end diff --git a/src/algebraic_geometry/morphisms/quasi_separated.lean b/src/algebraic_geometry/morphisms/quasi_separated.lean index f481d10e88d38..5ccde009c11cf 100644 --- a/src/algebraic_geometry/morphisms/quasi_separated.lean +++ b/src/algebraic_geometry/morphisms/quasi_separated.lean @@ -340,11 +340,9 @@ begin { simp only [X.basic_open_res], rintros x ⟨H₁, H₂⟩, exact ⟨h₂ H₁, H₂⟩ } } }, use n, - conv_lhs at e { rw mul_comm }, - conv_rhs at e { rw mul_comm }, simp only [pow_add, map_pow, map_mul, ← comp_apply, ← mul_assoc, ← functor.map_comp, subtype.coe_mk] at e ⊢, - convert e + exact e end lemma exists_eq_pow_mul_of_is_compact_of_is_quasi_separated (X : Scheme) @@ -468,15 +466,15 @@ begin using e.symm }, { intros x y, rw [← sub_eq_zero, ← map_sub, ring_hom.algebra_map_to_algebra], - simp_rw [← @sub_eq_zero _ _ (x * _) (y * _), ← sub_mul], + simp_rw [← @sub_eq_zero _ _ (_ * x) (_ * y), ← mul_sub], generalize : x - y = z, split, { intro H, obtain ⟨n, e⟩ := exists_pow_mul_eq_zero_of_res_basic_open_eq_zero_of_is_compact X hU _ _ H, refine ⟨⟨_, n, rfl⟩, _⟩, simpa [mul_comm z] using e }, - { rintro ⟨⟨_, n, rfl⟩, e : z * f ^ n = 0⟩, - rw [← ((RingedSpace.is_unit_res_basic_open _ f).pow n).mul_left_inj, zero_mul, ← map_pow, + { rintro ⟨⟨_, n, rfl⟩, e : f ^ n * z = 0⟩, + rw [← ((RingedSpace.is_unit_res_basic_open _ f).pow n).mul_right_inj, mul_zero, ← map_pow, ← map_mul, e, map_zero] } } end diff --git a/src/algebraic_geometry/projective_spectrum/scheme.lean b/src/algebraic_geometry/projective_spectrum/scheme.lean index 3d28c47c871a3..78fb0a006f8d7 100644 --- a/src/algebraic_geometry/projective_spectrum/scheme.lean +++ b/src/algebraic_geometry/projective_spectrum/scheme.lean @@ -191,11 +191,11 @@ begin change localization.mk (f ^ N) 1 = mk (∑ _, _) 1 at eq1, simp only [mk_eq_mk', is_localization.eq] at eq1, rcases eq1 with ⟨⟨_, ⟨M, rfl⟩⟩, eq1⟩, - erw [mul_one, mul_one] at eq1, - change f^_ * f^_ = _ * f^_ at eq1, + erw [one_mul, one_mul] at eq1, + change f^_ * f^_ = f^_ * _ at eq1, rw set.not_disjoint_iff_nonempty_inter, - refine ⟨f^N * f^M, eq1.symm ▸ mul_mem_right _ _ - (sum_mem _ (λ i hi, mul_mem_left _ _ _)), ⟨N+M, by rw pow_add⟩⟩, + refine ⟨f^M * f^N, eq1.symm ▸ mul_mem_left _ _ + (sum_mem _ (λ i hi, mul_mem_left _ _ _)), ⟨M + N, by rw pow_add⟩⟩, generalize_proofs h₁ h₂, exact (classical.some_spec h₂).1, end @@ -220,9 +220,8 @@ def to_fun (x : Proj.T| (pbo f)) : (Spec.T (A⁰_ f)) := simp only [localization.mk_mul, one_mul] at eq1, simp only [mk_eq_mk', is_localization.eq] at eq1, rcases eq1 with ⟨⟨_, ⟨M, rfl⟩⟩, eq1⟩, - rw [submonoid.coe_one, mul_one] at eq1, - change _ * _ * f^_ = _ * (f^_ * f^_) * f^_ at eq1, - + rw [submonoid.coe_one, one_mul] at eq1, + change f^_ * (_ * _) = f^_ * (f^_ * f^_ * _) at eq1, rcases x.1.is_prime.mem_or_mem (show a1 * a2 * f ^ N * f ^ M ∈ _, from _) with h1|rid2, rcases x.1.is_prime.mem_or_mem h1 with h1|rid1, rcases x.1.is_prime.mem_or_mem h1 with h1|h2, @@ -234,8 +233,8 @@ def to_fun (x : Proj.T| (pbo f)) : (Spec.T (A⁰_ f)) := exact ideal.mul_mem_right _ _ (ideal.subset_span ⟨_, h2, rfl⟩), }, { exact false.elim (x.2 (x.1.is_prime.mem_of_pow_mem N rid1)), }, { exact false.elim (x.2 (x.1.is_prime.mem_of_pow_mem M rid2)), }, - { rw [mul_comm _ (f^N), eq1], - refine mul_mem_right _ _ (mul_mem_right _ _ (sum_mem _ (λ i hi, mul_mem_left _ _ _))), + { rw [←mul_comm (f^M), ←mul_comm (f^N), eq1], + refine mul_mem_left _ _ (mul_mem_left _ _ (sum_mem _ (λ i hi, mul_mem_left _ _ _))), generalize_proofs h₁ h₂, exact (classical.some_spec h₂).1 }, end⟩ @@ -273,7 +272,7 @@ begin change localization.mk (f^N) 1 * mk _ _ = mk (∑ _, _) _ at eq1, rw [mk_mul, one_mul, mk_eq_mk', is_localization.eq] at eq1, rcases eq1 with ⟨⟨_, ⟨M, rfl⟩⟩, eq1⟩, - rw [submonoid.coe_one, mul_one] at eq1, + rw [submonoid.coe_one, one_mul] at eq1, simp only [subtype.coe_mk] at eq1, rcases y.1.is_prime.mem_or_mem (show a * f ^ N * f ^ M ∈ _, from _) with H1 | H3, @@ -281,8 +280,8 @@ begin { exact hy2 H1, }, { exact y.2 (y.1.is_prime.mem_of_pow_mem N H2), }, { exact y.2 (y.1.is_prime.mem_of_pow_mem M H3), }, - { rw [mul_comm _ (f^N), eq1], - refine mul_mem_right _ _ (mul_mem_right _ _ (sum_mem _ (λ i hi, mul_mem_left _ _ _))), + { rw [mul_comm _ (f^N), mul_comm _ (f^M), eq1], + refine mul_mem_left _ _ (mul_mem_left _ _ (sum_mem _ (λ i hi, mul_mem_left _ _ _))), generalize_proofs h₁ h₂, exact (classical.some_spec h₂).1, }, }, end diff --git a/src/algebraic_geometry/structure_sheaf.lean b/src/algebraic_geometry/structure_sheaf.lean index 236af107c5bda..c9dd40eec9e22 100644 --- a/src/algebraic_geometry/structure_sheaf.lean +++ b/src/algebraic_geometry/structure_sheaf.lean @@ -358,7 +358,8 @@ by convert is_localization.mk'_mul _ f₁ f₂ ⟨g₁, hu₁ x x.2⟩ ⟨g₂, lemma const_ext {f₁ f₂ g₁ g₂ : R} {U hu₁ hu₂} (h : f₁ * g₂ = f₂ * g₁) : const R f₁ g₁ U hu₁ = const R f₂ g₂ U hu₂ := -subtype.eq $ funext $ λ x, is_localization.mk'_eq_of_eq h.symm +subtype.eq $ funext $ λ x, is_localization.mk'_eq_of_eq + (by rw [mul_comm, subtype.coe_mk, ←h, mul_comm, subtype.coe_mk]) lemma const_congr {f₁ f₂ g₁ g₂ : R} {U hu} (hf : f₁ = f₂) (hg : g₁ = g₂) : const R f₁ g₁ U hu = const R f₂ g₂ U (hg ▸ hu) := @@ -575,17 +576,17 @@ begin rw is_localization.eq, -- We know that the fractions `a/b` and `c/d` are equal as sections of the structure sheaf on -- `basic_open f`. We need to show that they agree as elements in the localization of `R` at `f`. - -- This amounts showing that `a * d * r = c * b * r`, for some power `r = f ^ n` of `f`. + -- This amounts showing that `r * (d * a) = r * (b * c)`, for some power `r = f ^ n` of `f`. -- We define `I` as the ideal of *all* elements `r` satisfying the above equation. let I : ideal R := - { carrier := {r : R | a * d * r = c * b * r}, - zero_mem' := by simp only [set.mem_set_of_eq, mul_zero], - add_mem' := λ r₁ r₂ hr₁ hr₂, by { dsimp at hr₁ hr₂ ⊢, simp only [mul_add, hr₁, hr₂] }, - smul_mem' := λ r₁ r₂ hr₂, by { dsimp at hr₂ ⊢, simp only [mul_comm r₁ r₂, ← mul_assoc, hr₂] }}, + { carrier := {r : R | r * (d * a) = r * (b * c)}, + zero_mem' := by simp only [set.mem_set_of_eq, zero_mul], + add_mem' := λ r₁ r₂ hr₁ hr₂, by { dsimp at hr₁ hr₂ ⊢, simp only [add_mul, hr₁, hr₂] }, + smul_mem' := λ r₁ r₂ hr₂, by { dsimp at hr₂ ⊢, simp only [mul_assoc, hr₂] }}, -- Our claim now reduces to showing that `f` is contained in the radical of `I` suffices : f ∈ I.radical, { cases this with n hn, - exact ⟨⟨f ^ n, n, rfl⟩, hn⟩ }, + exact ⟨⟨f ^ n, n, rfl⟩, hn⟩, }, rw [← vanishing_ideal_zero_locus_eq_radical, mem_vanishing_ideal], intros p hfp, contrapose hfp, diff --git a/src/data/polynomial/laurent.lean b/src/data/polynomial/laurent.lean index 11243e114a4ae..bcdfada8689be 100644 --- a/src/data/polynomial/laurent.lean +++ b/src/data/polynomial/laurent.lean @@ -508,7 +508,7 @@ lemma is_localization : is_localization (submonoid.closure ({X} : set R[X])) R[T exact ⟨1, rfl⟩ }, { rintro ⟨⟨h, hX⟩, h⟩, rcases submonoid.mem_closure_singleton.mp hX with ⟨n, rfl⟩, - exact mul_X_pow_injective n (by simpa only [X_pow_mul] using h) } + exact mul_X_pow_injective n h } end } end comm_semiring diff --git a/src/field_theory/ratfunc.lean b/src/field_theory/ratfunc.lean index a7966a588eebc..f15d5206944ae 100644 --- a/src/field_theory/ratfunc.lean +++ b/src/field_theory/ratfunc.lean @@ -125,7 +125,7 @@ lemma to_fraction_ring_injective : /-- Non-dependent recursion principle for `ratfunc K`: To construct a term of `P : Sort*` out of `x : ratfunc K`, it suffices to provide a constructor `f : Π (p q : K[X]), P` -and a proof that `f p q = f p' q'` for all `p q p' q'` such that `p * q' = p' * q` where +and a proof that `f p q = f p' q'` for all `p q p' q'` such that `q' * p = q * p'` where both `q` and `q'` are not zero divisors, stated as `q ∉ K[X]⁰`, `q' ∉ K[X]⁰`. If considering `K` as an integral domain, this is the same as saying that @@ -137,18 +137,18 @@ of `∀ {p q a : K[X]} (hq : q ≠ 0) (ha : a ≠ 0), f (a * p) (a * q) = f p q) -/ @[irreducible] protected def lift_on {P : Sort v} (x : ratfunc K) (f : ∀ (p q : K[X]), P) - (H : ∀ {p q p' q'} (hq : q ∈ K[X]⁰) (hq' : q' ∈ K[X]⁰), p * q' = p' * q → + (H : ∀ {p q p' q'} (hq : q ∈ K[X]⁰) (hq' : q' ∈ K[X]⁰), q' * p = q * p' → f p q = f p' q') : P := localization.lift_on (by exact to_fraction_ring x) -- Fix timeout by manipulating elaboration order (λ p q, f p q) (λ p p' q q' h, H q.2 q'.2 (let ⟨⟨c, hc⟩, mul_eq⟩ := (localization.r_iff_exists).mp h in - mul_cancel_right_coe_non_zero_divisor.mp mul_eq)) + mul_cancel_left_coe_non_zero_divisor.mp mul_eq)) lemma lift_on_of_fraction_ring_mk {P : Sort v} (n : K[X]) (d : K[X]⁰) (f : ∀ (p q : K[X]), P) - (H : ∀ {p q p' q'} (hq : q ∈ K[X]⁰) (hq' : q' ∈ K[X]⁰), p * q' = p' * q → + (H : ∀ {p q p' q'} (hq : q ∈ K[X]⁰) (hq' : q' ∈ K[X]⁰), q' * p = q * p' → f p q = f p' q') : ratfunc.lift_on (by exact of_fraction_ring (localization.mk n d)) f @H = f n d := begin @@ -199,13 +199,13 @@ by rw [← is_localization.mk'_one (fraction_ring K[X]) p, ← mk_coe_def, submo lemma mk_eq_mk {p q p' q' : K[X]} (hq : q ≠ 0) (hq' : q' ≠ 0) : ratfunc.mk p q = ratfunc.mk p' q' ↔ p * q' = p' * q := by rw [mk_def_of_ne _ hq, mk_def_of_ne _ hq', of_fraction_ring_injective.eq_iff, - is_localization.mk'_eq_iff_eq, set_like.coe_mk, set_like.coe_mk, + is_localization.mk'_eq_iff_eq', set_like.coe_mk, set_like.coe_mk, (is_fraction_ring.injective K[X] (fraction_ring K[X])).eq_iff] lemma lift_on_mk {P : Sort v} (p q : K[X]) (f : ∀ (p q : K[X]), P) (f0 : ∀ p, f p 0 = f 0 1) - (H' : ∀ {p q p' q'} (hq : q ≠ 0) (hq' : q' ≠ 0), p * q' = p' * q → f p q = f p' q') - (H : ∀ {p q p' q'} (hq : q ∈ K[X]⁰) (hq' : q' ∈ K[X]⁰), p * q' = p' * q → + (H' : ∀ {p q p' q'} (hq : q ≠ 0) (hq' : q' ≠ 0), q' * p = q * p' → f p q = f p' q') + (H : ∀ {p q p' q'} (hq : q ∈ K[X]⁰) (hq' : q' ∈ K[X]⁰), q' * p = q * p' → f p q = f p' q' := λ p q p' q' hq hq' h, H' (non_zero_divisors.ne_zero hq) (non_zero_divisors.ne_zero hq') h) : (ratfunc.mk p q).lift_on f @H = f p q := @@ -218,25 +218,16 @@ begin set_like.coe_mk] } end +omit hdomain lemma lift_on_condition_of_lift_on'_condition {P : Sort v} {f : ∀ (p q : K[X]), P} (H : ∀ {p q a} (hq : q ≠ 0) (ha : a ≠ 0), f (a * p) (a * q) = f p q) - ⦃p q p' q' : K[X]⦄ (hq : q ≠ 0) (hq' : q' ≠ 0) (h : p * q' = p' * q) : + ⦃p q p' q' : K[X]⦄ (hq : q ≠ 0) (hq' : q' ≠ 0) (h : q' * p = q * p') : f p q = f p' q' := -begin - have H0 : f 0 q = f 0 q', - { calc f 0 q = f (q' * 0) (q' * q) : (H hq hq').symm - ... = f (q * 0) (q * q') : by rw [mul_zero, mul_zero, mul_comm] - ... = f 0 q' : H hq' hq }, - by_cases hp : p = 0, - { simp only [hp, hq, zero_mul, or_false, zero_eq_mul] at ⊢ h, rw [h, H0] }, - by_cases hp' : p' = 0, - { simpa only [hp, hp', hq', zero_mul, or_self, mul_eq_zero] using h }, - calc f p q = f (p' * p) (p' * q) : (H hq hp').symm - ... = f (p * p') (p * q') : by rw [mul_comm p p', h] - ... = f p' q' : H hq' hp -end +calc f p q = f (q' * p) (q' * q) : (H hq hq').symm + ... = f (q * p') (q * q') : by rw [h, mul_comm q'] + ... = f p' q' : H hq' hq +include hdomain --- f /-- Non-dependent recursion principle for `ratfunc K`: if `f p q : P` for all `p q`, such that `f (a * p) (a * q) = f p q`, then we can find a value of `P` for all elements of `ratfunc K` by setting `lift_on' (p / q) f _ = f p q`. @@ -496,7 +487,7 @@ def map [monoid_hom_class F R[X] S[X]] (φ : F) { exact hφ hq' }, { exact hφ hq }, refine localization.r_of_eq _, - simpa only [map_mul] using (congr_arg φ h).symm, + simpa only [map_mul] using (congr_arg φ h), end, map_one' := begin rw [←of_fraction_ring_one, ←localization.mk_one, lift_on_of_fraction_ring_mk, dif_pos], @@ -532,7 +523,7 @@ begin rintro ⟨x⟩ ⟨y⟩ h, induction x, induction y, { simpa only [map_apply_of_fraction_ring_mk, of_fraction_ring_injective.eq_iff, localization.mk_eq_mk_iff, localization.r_iff_exists, - mul_cancel_right_coe_non_zero_divisor, exists_const, set_like.coe_mk, ←map_mul, + mul_cancel_left_coe_non_zero_divisor, exists_const, set_like.coe_mk, ←map_mul, hf.eq_iff] using h }, { refl }, { refl } @@ -573,7 +564,7 @@ def lift_monoid_with_zero_hom (φ : R[X] →*₀ G₀) (hφ : R[X]⁰ ≤ G₀ { to_fun := λ f, ratfunc.lift_on f (λ p q, φ p / (φ q)) $ λ p q p' q' hq hq' h, begin casesI subsingleton_or_nontrivial R, { rw [subsingleton.elim p q, subsingleton.elim p' q, subsingleton.elim q' q] }, - rw [div_eq_div_iff, ←map_mul, h, map_mul]; + rw [div_eq_div_iff, ←map_mul, mul_comm p, h, map_mul, mul_comm]; exact non_zero_divisors.ne_zero (hφ ‹_›), end, map_one' := by { rw [←of_fraction_ring_one, ←localization.mk_one, lift_on_of_fraction_ring_mk], @@ -602,8 +593,9 @@ begin { simp_rw [lift_monoid_with_zero_hom_apply_of_fraction_ring_mk, localization.mk_eq_mk_iff], intro h, refine localization.r_of_eq _, - simpa only [←hφ.eq_iff, map_mul] using mul_eq_mul_of_div_eq_div _ _ _ _ h.symm; - exact (map_ne_zero_of_mem_non_zero_divisors _ hφ (set_like.coe_mem _)) }, + have := mul_eq_mul_of_div_eq_div _ _ _ _ h, + rwa [←map_mul, ←map_mul, hφ.eq_iff, mul_comm, mul_comm y_a] at this, + all_goals { exact (map_ne_zero_of_mem_non_zero_divisors _ hφ (set_like.coe_mem _)) } }, { exact λ _, rfl }, { exact λ _, rfl } end @@ -827,8 +819,8 @@ variables {K} @[simp] lemma lift_on_div {P : Sort v} (p q : K[X]) (f : ∀ (p q : K[X]), P) (f0 : ∀ p, f p 0 = f 0 1) - (H' : ∀ {p q p' q'} (hq : q ≠ 0) (hq' : q' ≠ 0), p * q' = p' * q → f p q = f p' q') - (H : ∀ {p q p' q'} (hq : q ∈ K[X]⁰) (hq' : q' ∈ K[X]⁰), p * q' = p' * q → + (H' : ∀ {p q p' q'} (hq : q ≠ 0) (hq' : q' ≠ 0), q' * p = q * p' → f p q = f p' q') + (H : ∀ {p q p' q'} (hq : q ∈ K[X]⁰) (hq' : q' ∈ K[X]⁰), q' * p = q * p' → f p q = f p' q' := λ p q p' q' hq hq' h, H' (non_zero_divisors.ne_zero hq) (non_zero_divisors.ne_zero hq') h) : (algebra_map _ (ratfunc K) p / algebra_map _ _ q).lift_on f @H = f p q := diff --git a/src/group_theory/monoid_localization.lean b/src/group_theory/monoid_localization.lean index b797ad903101e..2a89add0d748b 100644 --- a/src/group_theory/monoid_localization.lean +++ b/src/group_theory/monoid_localization.lean @@ -72,7 +72,7 @@ satisfies this predicate, then `N` is isomorphic to the localization of `M` at ` extends add_monoid_hom M N := (map_add_units' : ∀ y : S, is_add_unit (to_fun y)) (surj' : ∀ z : N, ∃ x : M × S, z + to_fun x.2 = to_fun x.1) -(eq_iff_exists' : ∀ x y, to_fun x = to_fun y ↔ ∃ c : S, x + c = y + c) +(eq_iff_exists' : ∀ x y, to_fun x = to_fun y ↔ ∃ c : S, ↑c + x = ↑c + y) /-- The add_monoid hom underlying a `localization_map` of `add_comm_monoid`s. -/ add_decl_doc localization_map.to_add_monoid_hom @@ -92,7 +92,7 @@ satisfies this predicate, then `N` is isomorphic to the localization of `M` at ` extends monoid_hom M N := (map_units' : ∀ y : S, is_unit (to_fun y)) (surj' : ∀ z : N, ∃ x : M × S, z * to_fun x.2 = to_fun x.1) -(eq_iff_exists' : ∀ x y, to_fun x = to_fun y ↔ ∃ c : S, x * c = y * c) +(eq_iff_exists' : ∀ x y, to_fun x = to_fun y ↔ ∃ c : S, ↑c * x = c * y) attribute [to_additive add_submonoid.localization_map] submonoid.localization_map attribute [to_additive add_submonoid.localization_map.to_add_monoid_hom] @@ -124,20 +124,21 @@ submonoid of `M`, whose quotient is the localization of `M` at `S`. -/ `S` a submonoid of `M`, whose quotient is the localization of `M` at `S`."] def r' : con (M × S) := begin - refine { r := λ a b : M × S, ∃ c : S, a.1 * b.2 * c = b.1 * a.2 * c, + -- note we multiply by `c` on the left so that we can later generalize to `•` + refine { r := λ a b : M × S, ∃ c : S, ↑c * (↑b.2 * a.1) = c * (a.2 * b.1), iseqv := ⟨λ a, ⟨1, rfl⟩, λ a b ⟨c, hc⟩, ⟨c, hc.symm⟩, _⟩, .. }, { rintros a b c ⟨t₁, ht₁⟩ ⟨t₂, ht₂⟩, - use b.2 * t₁ * t₂, + use t₂ * t₁ * b.2, simp only [submonoid.coe_mul], - calc a.1 * c.2 * (b.2 * t₁ * t₂) = a.1 * b.2 * t₁ * c.2 * t₂ : by ac_refl - ... = b.1 * c.2 * t₂ * a.2 * t₁ : by { rw ht₁, ac_refl } - ... = c.1 * a.2 * (b.2 * t₁ * t₂) : by { rw ht₂, ac_refl } }, + calc (t₂ * t₁ * b.2 : M) * (c.2 * a.1) = t₂ * c.2 * (t₁ * (b.2 * a.1)) : by ac_refl + ... = t₁ * a.2 * (t₂ * (c.2 * b.1)) : by { rw ht₁, ac_refl } + ... = t₂ * t₁ * b.2 * (a.2 * c.1) : by { rw ht₂, ac_refl } }, { rintros a b c d ⟨t₁, ht₁⟩ ⟨t₂, ht₂⟩, - use t₁ * t₂, - calc (a.1 * c.1) * (b.2 * d.2) * (t₁ * t₂) = (a.1 * b.2 * t₁) * (c.1 * d.2 * t₂) : + use t₂ * t₁, + calc (t₂ * t₁ : M) * ((b.2 * d.2) * (a.1 * c.1)) = (t₂ * (d.2 * c.1)) * (t₁ * (b.2 * a.1)) : by ac_refl - ... = (b.1 * d.1) * (a.2 * c.2) * (t₁ * t₂) : by { rw [ht₁, ht₂], ac_refl } } + ... = (t₂ * t₁ : M) * ((a.2 * c.2) * (b.1 * d.1)) : by { rw [ht₁, ht₂], ac_refl } } end /-- The congruence relation used to localize a `comm_monoid` at a submonoid can be expressed @@ -148,19 +149,19 @@ submonoid can be expressed equivalently as an infimum (see `add_localization.r`) explicitly (see `add_localization.r'`)."] theorem r_eq_r' : r S = r' S := le_antisymm (Inf_le $ λ _, ⟨1, by simp⟩) $ - le_Inf $ λ b H ⟨p, q⟩ y ⟨t, ht⟩, + le_Inf $ λ b H ⟨p, q⟩ ⟨x, y⟩ ⟨t, ht⟩, begin - rw [← mul_one (p, q), ← mul_one y], - refine b.trans (b.mul (b.refl _) (H (y.2 * t))) _, - convert b.symm (b.mul (b.refl y) (H (q * t))) using 1, - rw [prod.mk_mul_mk, submonoid.coe_mul, ← mul_assoc, ht, mul_left_comm, mul_assoc], - refl + rw [← one_mul (p, q), ← one_mul (x, y)], + refine b.trans (b.mul (H (t * y)) (b.refl _)) _, + convert b.symm (b.mul (H (t * q)) (b.refl (x, y))) using 1, + dsimp only [prod.mk_mul_mk, submonoid.coe_mul] at ⊢ ht, + simp_rw [mul_assoc, ht, mul_comm y q], end variables {S} @[to_additive] -lemma r_iff_exists {x y : M × S} : r S x y ↔ ∃ c : S, x.1 * y.2 * c = y.1 * x.2 * c := +lemma r_iff_exists {x y : M × S} : r S x y ↔ ∃ c : S, ↑c * (↑y.2 * x.1) = c * (x.2 * y.1) := by rw r_eq_r' S; refl end localization @@ -323,7 +324,7 @@ induction_on₂ x y $ λ x y, induction_on z $ H x y @[to_additive] lemma one_rel (y : S) : r S 1 (y, y) := λ b hb, hb y -@[to_additive] theorem r_of_eq {x y : M × S} (h : y.1 * x.2 = x.1 * y.2) : r S x y := +@[to_additive] theorem r_of_eq {x y : M × S} (h : ↑y.2 * x.1 = ↑x.2 * y.1) : r S x y := r_iff_exists.2 ⟨1, by rw h⟩ @[to_additive] lemma mk_self (a : S) : mk (a : M) a = 1 := @@ -344,7 +345,12 @@ begin rw r_eq_r' at h ⊢, cases h with t ht, use t, - simp only [smul_mul_assoc, ht] + dsimp only [subtype.coe_mk] at ht ⊢, + -- TODO: this definition should take `smul_comm_class R M M` instead of `is_scalar_tower R M M` if + -- we ever want to generalize to the non-commutative case. + haveI : smul_comm_class R M M := + ⟨λ r m₁ m₂, by simp_rw [smul_eq_mul, mul_comm m₁, smul_mul_assoc]⟩, + simp only [mul_smul_comm, ht], end instance [has_smul R M] [is_scalar_tower R M M] : @@ -406,7 +412,7 @@ namespace monoid_hom @[to_additive "Makes a localization map from an `add_comm_monoid` hom satisfying the characteristic predicate."] def to_localization_map (f : M →* N) (H1 : ∀ y : S, is_unit (f y)) - (H2 : ∀ z, ∃ x : M × S, z * f x.2 = f x.1) (H3 : ∀ x y, f x = f y ↔ ∃ c : S, x * c = y * c) : + (H2 : ∀ z, ∃ x : M × S, z * f x.2 = f x.1) (H3 : ∀ x y, f x = f y ↔ ∃ c : S, ↑c * x = ↑c * y) : submonoid.localization_map S N := { map_units' := H1, surj' := H2, @@ -440,7 +446,7 @@ by { rcases f with ⟨⟨⟩⟩, rcases g with ⟨⟨⟩⟩, simp only, exact fu ∃ x : M × S, z * f.to_map x.2 = f.to_map x.1 := f.3 z @[to_additive] lemma eq_iff_exists (f : localization_map S N) {x y} : - f.to_map x = f.to_map y ↔ ∃ c : S, x * c = y * c := f.4 x y + f.to_map x = f.to_map y ↔ ∃ c : S, ↑c * x = c * y := f.4 x y /-- Given a localization map `f : M →* N`, a section function sending `z : N` to some `(x, y) : M × S` such that `f x * (f y)⁻¹ = z`. -/ @@ -565,14 +571,18 @@ by rw [mul_comm, mk'_spec] by rw [eq_comm, eq_mk'_iff_mul_eq, eq_comm] @[to_additive] lemma mk'_eq_iff_eq {x₁ x₂} {y₁ y₂ : S} : - f.mk' x₁ y₁ = f.mk' x₂ y₂ ↔ f.to_map (x₁ * y₂) = f.to_map (x₂ * y₁) := -⟨λ H, by rw [f.to_map.map_mul, f.mk'_eq_iff_eq_mul.1 H, mul_assoc, - mul_comm (f.to_map _), ←mul_assoc, mk'_spec, f.to_map.map_mul], + f.mk' x₁ y₁ = f.mk' x₂ y₂ ↔ f.to_map (y₂ * x₁) = f.to_map (y₁ * x₂) := +⟨λ H, by rw [f.to_map.map_mul, f.to_map.map_mul, f.mk'_eq_iff_eq_mul.1 H, ←mul_assoc, mk'_spec', + mul_comm], λ H, by rw [mk'_eq_iff_eq_mul, mk', mul_assoc, mul_comm _ (f.to_map y₁), ←mul_assoc, - ←f.to_map.map_mul, ←H, f.to_map.map_mul, mul_inv_right f.map_units]⟩ + ←f.to_map.map_mul, mul_comm x₂, ←H, ←mul_comm x₁, f.to_map.map_mul, mul_inv_right f.map_units]⟩ + +@[to_additive] lemma mk'_eq_iff_eq' {x₁ x₂} {y₁ y₂ : S} : + f.mk' x₁ y₁ = f.mk' x₂ y₂ ↔ f.to_map (x₁ * y₂) = f.to_map (x₂ * y₁) := +by simp only [f.mk'_eq_iff_eq, mul_comm] @[to_additive] protected lemma eq {a₁ b₁} {a₂ b₂ : S} : - f.mk' a₁ a₂ = f.mk' b₁ b₂ ↔ ∃ c : S, a₁ * b₂ * c = b₁ * a₂ * c := + f.mk' a₁ a₂ = f.mk' b₁ b₂ ↔ ∃ c : S, ↑c * (↑b₂ * a₁) = c * (a₂ * b₁) := f.mk'_eq_iff_eq.trans $ f.eq_iff_exists @[to_additive] protected lemma eq' {a₁ b₁} {a₂ b₂ : S} : @@ -594,13 +604,17 @@ such that `x₁ * y₂ * c = x₂ * y₁ * c`. -/ and `y₁ ∈ S`, if `x₂ : M, y₂ ∈ S` are such that `(f x₁ - f y₁) + f y₂ = f x₂`, then there exists `c ∈ S` such that `x₁ + y₂ + c = x₂ + y₁ + c`."] lemma exists_of_sec_mk' (x) (y : S) : - ∃ c : S, x * (f.sec $ f.mk' x y).2 * c = (f.sec $ f.mk' x y).1 * y * c := + ∃ c : S, ↑c * (↑(f.sec $ f.mk' x y).2 * x) = c * (y * (f.sec $ f.mk' x y).1) := f.eq_iff_exists.1 $ f.mk'_eq_iff_eq.1 $ (mk'_sec _ _).symm -@[to_additive] lemma mk'_eq_of_eq {a₁ b₁ : M} {a₂ b₂ : S} (H : b₁ * a₂ = a₁ * b₂) : +@[to_additive] lemma mk'_eq_of_eq {a₁ b₁ : M} {a₂ b₂ : S} (H : ↑a₂ * b₁ = ↑b₂ * a₁) : f.mk' a₁ a₂ = f.mk' b₁ b₂ := f.mk'_eq_iff_eq.2 $ H ▸ rfl +@[to_additive] lemma mk'_eq_of_eq' {a₁ b₁ : M} {a₂ b₂ : S} (H : b₁ * ↑a₂ = a₁ * ↑b₂) : + f.mk' a₁ a₂ = f.mk' b₁ b₂ := +f.mk'_eq_of_eq $ by simpa only [mul_comm] using H + @[simp, to_additive] lemma mk'_self' (y : S) : f.mk' (y : M) y = 1 := show _ * _ = _, by rw [mul_inv_left, mul_one] @@ -646,9 +660,9 @@ lemma eq_of_eq (hg : ∀ y : S, is_unit (g y)) {x y} (h : f.to_map x = f.to_map g x = g y := begin obtain ⟨c, hc⟩ := f.eq_iff_exists.1 h, - rw [←mul_one (g x), ←is_unit.mul_lift_right_inv (g.restrict S) hg c], - show _ * (g c * _) = _, - rw [←mul_assoc, ←g.map_mul, hc, mul_inv_left hg, g.map_mul, mul_comm], + rw [←one_mul (g x), ←is_unit.lift_right_inv_mul (g.restrict S) hg c], + show (_ * g c) * _ = _, + rw [mul_assoc, ←g.map_mul, hc, mul_comm, mul_inv_left hg, g.map_mul], end /-- Given `comm_monoid`s `M, P`, localization maps `f : M →* N, k : P →* Q` for submonoids @@ -1212,7 +1226,7 @@ def monoid_of : submonoid.localization_map S (localization S) := surj' := λ z, induction_on z $ λ x, ⟨x, by rw [mk_mul, mul_comm x.fst, ← mk_mul, mk_self, one_mul]⟩, eq_iff_exists' := λ x y, mk_eq_mk_iff.trans $ r_iff_exists.trans $ - show (∃ (c : S), x * 1 * c = y * 1 * c) ↔ _, by rw [mul_one, mul_one], + show (∃ (c : S), ↑c * (1 * x) = c * (1 * y)) ↔ _, by rw [one_mul, one_mul], ..(r S).mk'.comp $ monoid_hom.inl M S } variables {S} diff --git a/src/number_theory/cyclotomic/basic.lean b/src/number_theory/cyclotomic/basic.lean index 652614ab0c3fc..966eba91a1f36 100644 --- a/src/number_theory/cyclotomic/basic.lean +++ b/src/number_theory/cyclotomic/basic.lean @@ -660,7 +660,7 @@ instance [ne_zero ((n : ℕ) : A)] : simp only [map_mul] } end, eq_iff_exists := λ x y, ⟨λ h, ⟨1, by rw adjoin_algebra_injective n A K h⟩, - λ ⟨c, hc⟩, by rw mul_right_cancel₀ (non_zero_divisors.ne_zero c.prop) hc⟩ } + λ ⟨c, hc⟩, by rw mul_left_cancel₀ (non_zero_divisors.ne_zero c.prop) hc⟩ } lemma eq_adjoin_primitive_root {μ : (cyclotomic_field n K)} (h : is_primitive_root μ n) : cyclotomic_ring n A K = adjoin A ({μ} : set ((cyclotomic_field n K))) := diff --git a/src/number_theory/padics/padic_integers.lean b/src/number_theory/padics/padic_integers.lean index b70cfcbe73956..89711aa75ff29 100644 --- a/src/number_theory/padics/padic_integers.lean +++ b/src/number_theory/padics/padic_integers.lean @@ -574,7 +574,7 @@ instance is_fraction_ring : is_fraction_ring ℤ_[p] ℚ_[p] := rw [algebra_map_apply, algebra_map_apply, subtype.coe_inj], refine ⟨λ h, ⟨1, by rw h⟩, _⟩, rintro ⟨⟨c, hc⟩, h⟩, - exact (mul_eq_mul_right_iff.mp h).resolve_right (mem_non_zero_divisors_iff_ne_zero.mp hc) + exact (mul_eq_mul_left_iff.mp h).resolve_right (mem_non_zero_divisors_iff_ne_zero.mp hc) end } end fraction_ring diff --git a/src/ring_theory/artinian.lean b/src/ring_theory/artinian.lean index 4ca5581b77375..e720b4146490f 100644 --- a/src/ring_theory/artinian.lean +++ b/src/ring_theory/artinian.lean @@ -456,8 +456,8 @@ begin rw [smul_eq_mul, smul_eq_mul, pow_succ', mul_assoc] at hr, apply_fun algebra_map R L at hr, simp only [map_mul, ←submonoid.coe_pow] at hr, - rw [←is_localization.mk'_one L, is_localization.mk'_eq_iff_eq, one_mul, submonoid.coe_one, - ←(is_localization.map_units L (s ^ n)).mul_left_cancel hr, map_mul, mul_comm], + rw [←is_localization.mk'_one L, is_localization.mk'_eq_iff_eq, mul_one, submonoid.coe_one, + ←(is_localization.map_units L (s ^ n)).mul_left_cancel hr, map_mul], end lemma localization_artinian : is_artinian_ring L := diff --git a/src/ring_theory/graded_algebra/homogeneous_localization.lean b/src/ring_theory/graded_algebra/homogeneous_localization.lean index 714d3f8767d65..f311442ab436f 100644 --- a/src/ring_theory/graded_algebra/homogeneous_localization.lean +++ b/src/ring_theory/graded_algebra/homogeneous_localization.lean @@ -477,14 +477,14 @@ lemma is_unit_iff_is_unit_val (f : homogeneous_localization.at_prime 𝒜 𝔭) localization.mk_eq_mk', is_localization.eq] at eq1, rcases eq1 with ⟨⟨c, hc⟩, eq1⟩, simp only [← subtype.val_eq_coe] at eq1, - change a * f.num * 1 * c = _ at eq1, + change c * (1 * (a * f.num)) = _ at eq1, simp only [one_mul, mul_one] at eq1, - have mem1 : a * f.num * c ∈ 𝔭.prime_compl := - eq1.symm ▸ λ r, or.elim (ideal.is_prime.mem_or_mem infer_instance r) (by tauto)(by tauto), + have mem1 : c * (a * f.num) ∈ 𝔭.prime_compl := + eq1.symm ▸ λ r, or.elim (ideal.is_prime.mem_or_mem infer_instance r) (by tauto) (by tauto), have mem2 : f.num ∉ 𝔭, { contrapose! mem1, erw [not_not], - exact ideal.mul_mem_right _ _ (ideal.mul_mem_left _ _ mem1), }, + exact ideal.mul_mem_left _ _ (ideal.mul_mem_left _ _ mem1), }, refine ⟨⟨f, quotient.mk' ⟨f.deg, ⟨f.denom, f.denom_mem_deg⟩, ⟨f.num, f.num_mem_deg⟩, mem2⟩, _, _⟩, rfl⟩; simp only [ext_iff_val, mul_val, val_mk', ← subtype.val_eq_coe, f.eq_num_div_denom, diff --git a/src/ring_theory/laurent_series.lean b/src/ring_theory/laurent_series.lean index 99ba18b74afa7..137a2f03ff5a7 100644 --- a/src/ring_theory/laurent_series.lean +++ b/src/ring_theory/laurent_series.lean @@ -135,14 +135,13 @@ rfl { rintro rfl, exact ⟨1, rfl⟩ }, { rintro ⟨⟨_, n, rfl⟩, hc⟩, - rw [← sub_eq_zero, ← sub_mul, power_series.ext_iff] at hc, + rw [← sub_eq_zero, ← mul_sub, power_series.ext_iff] at hc, rw [← sub_eq_zero, power_series.ext_iff], intro m, have h := hc (m + n), - rw [linear_map.map_zero, subtype.coe_mk, power_series.X_pow_eq, power_series.monomial, - power_series.coeff, finsupp.single_add, mv_power_series.coeff_add_mul_monomial, - mul_one] at h, - exact h } end) } + rwa [linear_map.map_zero, subtype.coe_mk, power_series.X_pow_eq, power_series.monomial, + add_comm m, power_series.coeff, finsupp.single_add, mv_power_series.coeff_add_monomial_mul, + one_mul] at h } end) } instance {K : Type u} [field K] : is_fraction_ring (power_series K) (laurent_series K) := is_localization.of_le (submonoid.powers (power_series.X : power_series K)) _ diff --git a/src/ring_theory/local_properties.lean b/src/ring_theory/local_properties.lean index 1e60de6874cf7..70b0e5931f852 100644 --- a/src/ring_theory/local_properties.lean +++ b/src/ring_theory/local_properties.lean @@ -244,8 +244,8 @@ begin rw [← _root_.map_mul, ← sub_eq_zero, ← map_sub] at eq, obtain ⟨⟨m, hm⟩, eq⟩ := (is_localization.map_eq_zero_iff P.prime_compl _ _).mp eq, refine hs ((hP.is_prime.mem_or_mem (le (ideal.mem_colon_singleton.mpr _))).resolve_right hm), - simp only [subtype.coe_mk, sub_mul, sub_eq_zero, mul_assoc] at eq, - simpa only [eq, mul_comm] using J.mul_mem_right m ha + simp only [subtype.coe_mk, mul_sub, sub_eq_zero, mul_comm x s, mul_left_comm] at eq, + simpa only [mul_assoc, eq] using J.mul_mem_left m ha end /-- Let `I J : ideal R`. If the localization of `I` at each maximal ideal `P` is equal to @@ -307,8 +307,8 @@ begin rw [← (algebra_map R S).map_zero] at hx', obtain ⟨m', hm'⟩ := (is_localization.eq_iff_exists M S).mp hx', apply_fun (*m'^n) at hm', - simp only [mul_assoc, zero_mul] at hm', - rw [mul_comm, ← pow_succ, ← mul_pow] at hm', + simp only [mul_assoc, zero_mul, mul_zero] at hm', + rw [← mul_left_comm, ← pow_succ, ← mul_pow] at hm', replace hm' := is_nilpotent.eq_zero ⟨_, hm'.symm⟩, rw [← (is_localization.map_units S m).mul_left_inj, hx, zero_mul, is_localization.map_eq_zero_iff M], @@ -361,8 +361,8 @@ begin erw is_localization.map_mk' at e', rw [eq_comm, is_localization.eq_mk'_iff_mul_eq, subtype.coe_mk, subtype.coe_mk, ← map_mul] at e', obtain ⟨⟨_, n', rfl⟩, e''⟩ := (is_localization.eq_iff_exists (submonoid.powers (f r)) _).mp e', - rw [subtype.coe_mk, mul_assoc, ← map_pow, ← map_mul, ← map_mul, ← pow_add, mul_comm] at e'', - exact ⟨n + n', _, e''.symm⟩ + rw [subtype.coe_mk, mul_comm x, ←mul_assoc, ← map_pow, ← map_mul, ← map_mul, ← pow_add] at e'', + exact ⟨n' + n, _, e''.symm⟩ end end surjective @@ -454,9 +454,9 @@ begin convert (submodule.span R (is_localization.finset_integer_multiple (submonoid.map (algebra_map R S) M) s : set S)).smul_mem a hx' using 1, convert ha₂.symm, - { rw [mul_comm (y' • x), subtype.coe_mk, submonoid.smul_def, submonoid.coe_mul, ← smul_smul], + { rw [subtype.coe_mk, submonoid.smul_def, submonoid.coe_mul, ← smul_smul], exact algebra.smul_def _ _ }, - { rw mul_comm, exact algebra.smul_def _ _ } + { exact algebra.smul_def _ _ } end /-- If `S` is an `R' = M⁻¹R` algebra, and `x ∈ span R' s`, @@ -613,10 +613,9 @@ begin rw [algebra.smul_def, ← _root_.map_mul] at hx'', obtain ⟨a, ha₂⟩ := (is_localization.eq_iff_exists M S').mp hx'', use a * y ^ n, - convert A.mul_mem hx' (hA₂ a.2), - convert ha₂.symm, - simp only [submonoid.smul_def, submonoid.coe_pow, smul_eq_mul, submonoid.coe_mul], - ring, + convert A.mul_mem hx' (hA₂ a.prop), + rw [submonoid.smul_def, smul_eq_mul, submonoid.coe_mul, submonoid.coe_pow, mul_assoc, ←ha₂, + mul_comm], end /-- diff --git a/src/ring_theory/localization/at_prime.lean b/src/ring_theory/localization/at_prime.lean index 62cd3aaab6ebd..f37c5f099febf 100644 --- a/src/ring_theory/localization/at_prime.lean +++ b/src/ring_theory/localization/at_prime.lean @@ -89,12 +89,12 @@ begin rw ←hrx at hx, rw ←hry at hy, obtain ⟨t, ht⟩ := is_localization.eq.1 hxyz, simp only [mul_one, one_mul, submonoid.coe_mul, subtype.coe_mk] at ht, - suffices : ↑sx * ↑sy * ↑sz * ↑t ∈ I, from + suffices : ↑t * (↑sx * ↑sy * ↑sz) ∈ I, from not_or (mt hp.mem_or_mem $ not_or sx.2 sy.2) sz.2 - (hp.mem_or_mem $ (hp.mem_or_mem this).resolve_right t.2), - rw [←ht, mul_assoc], - exact I.mul_mem_right _ (I.add_mem (I.mul_mem_right _ $ this hx) - (I.mul_mem_right _ $ this hy)) + (hp.mem_or_mem $ (hp.mem_or_mem this).resolve_left t.2), + rw [←ht], + exact I.mul_mem_left _ (I.mul_mem_right _ (I.add_mem (I.mul_mem_right _ $ this hx) + (I.mul_mem_right _ $ this hy))), end end is_localization diff --git a/src/ring_theory/localization/away.lean b/src/ring_theory/localization/away.lean index 1ca217eb416a1..39d04cb19893d 100644 --- a/src/ring_theory/localization/away.lean +++ b/src/ring_theory/localization/away.lean @@ -104,7 +104,7 @@ begin { intros x y hxy, obtain ⟨c, eq⟩ := (is_localization.eq_iff_exists M S).mp hxy, obtain ⟨u, hu⟩ := H c, - rwa [← hu, units.mul_left_inj] at eq }, + rwa [← hu, units.mul_right_inj] at eq }, { intros y, obtain ⟨⟨x, s⟩, eq⟩ := is_localization.surj M y, obtain ⟨u, hu⟩ := H s, @@ -140,7 +140,7 @@ lemma away_of_is_unit_of_bijective {R : Type*} (S : Type*) [comm_ring R] [comm_r erw H.1.eq_iff, split, { rintro rfl, exact ⟨1, rfl⟩ }, - { rintro ⟨⟨_, n, rfl⟩, e⟩, exact (hr.pow _).mul_left_inj.mp e } + { rintro ⟨⟨_, n, rfl⟩, e⟩, exact (hr.pow _).mul_right_inj.mp e } end } end at_units diff --git a/src/ring_theory/localization/basic.lean b/src/ring_theory/localization/basic.lean index aff3a6372c8cc..63038c1b0b0f0 100644 --- a/src/ring_theory/localization/basic.lean +++ b/src/ring_theory/localization/basic.lean @@ -93,7 +93,7 @@ expresses that `S` is isomorphic to the localization of `R` at `M`. -/ class is_localization : Prop := (map_units [] : ∀ y : M, is_unit (algebra_map R S y)) (surj [] : ∀ z : S, ∃ x : R × M, z * algebra_map R S x.2 = algebra_map R S x.1) -(eq_iff_exists [] : ∀ {x y}, algebra_map R S x = algebra_map R S y ↔ ∃ c : M, x * c = y * c) +(eq_iff_exists [] : ∀ {x y}, algebra_map R S x = algebra_map R S y ↔ ∃ c : M, ↑c * x = ↑c * y) variables {M S} @@ -117,7 +117,7 @@ lemma of_le (N : submonoid R) (h₁ : M ≤ N) rintro ⟨c, hc⟩, exact ⟨⟨c, h₁ c.2⟩, hc⟩ }, { rintro ⟨c, h⟩, - simpa only [set_like.coe_mk, map_mul, (h₂ c c.2).mul_left_inj] using + simpa only [set_like.coe_mk, map_mul, (h₂ c c.2).mul_right_inj] using congr_arg (algebra_map R S) h } end } @@ -187,7 +187,7 @@ by { rw [hx, (algebra_map R S).map_zero] at h, variables (M S) lemma map_eq_zero_iff (r : R) : - algebra_map R S r = 0 ↔ ∃ m : M, r * m = 0 := + algebra_map R S r = 0 ↔ ∃ m : M, ↑m * r = 0 := begin split, intro h, @@ -195,7 +195,7 @@ begin ((algebra_map R S).map_zero.trans h.symm), exact ⟨m, by simpa using hm.symm⟩ }, { rintro ⟨m, hm⟩, - rw [← (is_localization.map_units S m).mul_left_inj, zero_mul, ← ring_hom.map_mul, hm, + rw [← (is_localization.map_units S m).mul_right_inj, mul_zero, ← ring_hom.map_mul, hm, ring_hom.map_zero] } end @@ -275,9 +275,13 @@ def unique_of_zero_mem (h : (0 : R) ∈ M) : unique S := unique_of_zero_eq_one $ by simpa using is_localization.map_units S ⟨0, h⟩ lemma mk'_eq_iff_eq {x₁ x₂} {y₁ y₂ : M} : - mk' S x₁ y₁ = mk' S x₂ y₂ ↔ algebra_map R S (x₁ * y₂) = algebra_map R S (x₂ * y₁) := + mk' S x₁ y₁ = mk' S x₂ y₂ ↔ algebra_map R S (y₂ * x₁) = algebra_map R S (y₁ * x₂) := (to_localization_map M S).mk'_eq_iff_eq +lemma mk'_eq_iff_eq' {x₁ x₂} {y₁ y₂ : M} : + mk' S x₁ y₁ = mk' S x₂ y₂ ↔ algebra_map R S (x₁ * y₂) = algebra_map R S (x₂ * y₁) := +(to_localization_map M S).mk'_eq_iff_eq' + lemma mk'_mem_iff {x} {y : M} {I : ideal S} : mk' S x y ∈ I ↔ algebra_map R S x ∈ I := begin split; @@ -291,11 +295,11 @@ begin end protected lemma eq {a₁ b₁} {a₂ b₂ : M} : - mk' S a₁ a₂ = mk' S b₁ b₂ ↔ ∃ c : M, a₁ * b₂ * c = b₁ * a₂ * c := + mk' S a₁ a₂ = mk' S b₁ b₂ ↔ ∃ c : M, ↑c * (↑b₂ * a₁) = c * (a₂ * b₁) := (to_localization_map M S).eq lemma mk'_eq_zero_iff (x : R) (s : M) : - mk' S x s = 0 ↔ ∃ (m : M), x * m = 0 := + mk' S x s = 0 ↔ ∃ (m : M), ↑m * x = 0 := by rw [← (map_units S s).mul_left_inj, mk'_spec, zero_mul, map_eq_zero_iff M] @[simp] lemma mk'_zero (s : M) : is_localization.mk' S 0 s = 0 := @@ -319,10 +323,14 @@ lemma mk'_eq_iff_mk'_eq {x₁ x₂} {y₁ y₂ : M} : mk' S x₁ y₁ = mk' S x₂ y₂ ↔ mk' P x₁ y₁ = mk' P x₂ y₂ := (to_localization_map M S).mk'_eq_iff_mk'_eq (to_localization_map M P) -lemma mk'_eq_of_eq {a₁ b₁ : R} {a₂ b₂ : M} (H : b₁ * a₂ = a₁ * b₂) : +lemma mk'_eq_of_eq {a₁ b₁ : R} {a₂ b₂ : M} (H : ↑a₂ * b₁ = ↑b₂ * a₁) : mk' S a₁ a₂ = mk' S b₁ b₂ := (to_localization_map M S).mk'_eq_of_eq H +lemma mk'_eq_of_eq' {a₁ b₁ : R} {a₂ b₂ : M} (H : b₁ * ↑a₂ = a₁ * ↑b₂) : + mk' S a₁ a₂ = mk' S b₁ b₂ := +(to_localization_map M S).mk'_eq_of_eq' H + variables (S) @[simp] lemma mk'_self {x : R} (hx : x ∈ M) : mk' S x ⟨x, hx⟩ = 1 := @@ -681,7 +689,7 @@ begin rw [ring_hom.algebra_map_to_algebra, ring_hom.comp_apply, ring_hom.comp_apply, is_localization.eq_iff_exists M S], simp_rw ← h.to_equiv.apply_eq_iff_eq, - change (∃ (c : M), h (h.symm x * c) = h (h.symm y * c)) ↔ _, + change (∃ (c : M), h (c * h.symm x) = h (c * h.symm y)) ↔ _, simp only [ring_equiv.apply_symm_apply, ring_equiv.map_mul], exact ⟨λ ⟨c, e⟩, ⟨⟨_, _, c.prop, rfl⟩, e⟩, λ ⟨⟨_, c, h, e₁⟩, e₂⟩, ⟨⟨_, h⟩, e₁.symm ▸ e₂⟩⟩ } end @@ -717,7 +725,7 @@ begin rw [← @mk'_one R _ M, ← mk'_mul, ← (algebra_map R S).map_zero, ← @mk'_one R _ M, is_localization.eq] at e, obtain ⟨c, e⟩ := e, - rw [zero_mul, zero_mul, submonoid.coe_one, mul_one, mul_comm x a, mul_assoc, mul_comm] at e, + rw [mul_zero, mul_zero, submonoid.coe_one, one_mul, ←mul_assoc] at e, rw mk'_eq_zero_iff, exact ⟨c, ha _ e⟩ end @@ -754,10 +762,11 @@ begin rw r_eq_r' at h1 h2 ⊢, cases h1 with t₅ ht₅, cases h2 with t₆ ht₆, - use t₆ * t₅, - calc ((b : R) * c + d * a) * (b' * d') * (t₆ * t₅) = - (c * d' * t₆) * (b * b' * t₅) + (a * b' * t₅) * (d * d' * t₆) : by ring - ... = (b' * c' + d' * a') * (b * d) * (t₆ * t₅) : by rw [ht₆, ht₅]; ring + use t₅ * t₆, + dsimp only, + calc (↑t₅ * ↑t₆) * ((↑b' * ↑d') * ((b : R) * c + d * a)) = + (t₆ * (d' * c)) * (t₅ * (b' * b)) + (t₅ * (b' * a)) * (t₆ * (d' * d)) : by ring + ... = (t₅ * t₆) * ((b * d) * (b' * c' + d' * a')) : by rw [ht₆, ht₅]; ring end instance : has_add (localization M) := ⟨localization.add⟩ @@ -941,7 +950,7 @@ begin rw r_eq_r' at h ⊢, cases h with t ht, use t, - rw [neg_mul, neg_mul, ht], + rw [mul_neg, mul_neg, ht], ring_nf, end @@ -993,7 +1002,7 @@ begin rw ← (algebra_map R S).map_zero, split; intro h, { cases (eq_iff_exists M S).mp h with c hc, - rw zero_mul at hc, + rw [mul_zero, mul_comm] at hc, exact hM c.2 x hc }, { rw h }, end diff --git a/src/ring_theory/localization/fraction_ring.lean b/src/ring_theory/localization/fraction_ring.lean index f726dc02c6f05..d7d0d63870b5a 100644 --- a/src/ring_theory/localization/fraction_ring.lean +++ b/src/ring_theory/localization/fraction_ring.lean @@ -57,7 +57,7 @@ instance rat.is_fraction_ring : is_fraction_ring ℤ ℚ := rw [eq_int_cast, eq_int_cast, int.cast_inj], refine ⟨by { rintro rfl, use 1 }, _⟩, rintro ⟨⟨c, hc⟩, h⟩, - apply mul_right_cancel₀ _ h, + apply mul_left_cancel₀ _ h, rwa mem_non_zero_divisors_iff_ne_zero at hc, end } @@ -165,7 +165,7 @@ is_unit.mk0 (g y) $ show g.to_monoid_with_zero_hom y ≠ 0, {y : non_zero_divisors R} : mk' K x y = 0 ↔ x = 0 := begin refine ⟨λ hxy, _, λ h, by rw [h, mk'_zero]⟩, - { simp_rw [mk'_eq_zero_iff, mul_right_coe_non_zero_divisors_eq_zero_iff] at hxy, + { simp_rw [mk'_eq_zero_iff, mul_left_coe_non_zero_divisors_eq_zero_iff] at hxy, exact (exists_const _).mp hxy }, end diff --git a/src/ring_theory/localization/ideal.lean b/src/ring_theory/localization/ideal.lean index 3d018025a297e..2fdb7a3b9f348 100644 --- a/src/ring_theory/localization/ideal.lean +++ b/src/ring_theory/localization/ideal.lean @@ -80,10 +80,11 @@ theorem comap_map_of_is_prime_disjoint (I : ideal R) (hI : I.is_prime) begin refine le_antisymm (λ a ha, _) ideal.le_comap_map, obtain ⟨⟨b, s⟩, h⟩ := (mem_map_algebra_map_iff M S).1 (ideal.mem_comap.1 ha), - replace h : algebra_map R S (a * s) = algebra_map R S b := by simpa only [←map_mul] using h, + replace h : algebra_map R S (s * a) = algebra_map R S b := + by simpa only [←map_mul, mul_comm] using h, obtain ⟨c, hc⟩ := (eq_iff_exists M S).1 h, - have : a * (s * c) ∈ I := by { rw [←mul_assoc, hc], exact I.mul_mem_right c b.2 }, - exact (hI.mem_or_mem this).resolve_right (λ hsc, hM.le_bot ⟨(s * c).2, hsc⟩) + have : (↑c * ↑s) * a ∈ I := by { rw [mul_assoc, hc], exact I.mul_mem_left c b.2 }, + exact (hI.mem_or_mem this).resolve_left (λ hsc, hM.le_bot ⟨(c * s).2, hsc⟩) end /-- If `S` is the localization of `R` at a submonoid, the ordering of ideals of `S` is diff --git a/src/ring_theory/localization/integral.lean b/src/ring_theory/localization/integral.lean index f6bdd17f8d5b0..0f9d2fd5a2851 100644 --- a/src/ring_theory/localization/integral.lean +++ b/src/ring_theory/localization/integral.lean @@ -309,7 +309,7 @@ lemma is_fraction_ring_of_algebraic (alg : is_algebraic A L) (by rw [is_scalar_tower.algebra_map_apply A C L, h, ring_hom.map_zero])))⟩, by rw [set_like.coe_mk, algebra_map_mk', ← is_scalar_tower.algebra_map_apply A C L, hxy]⟩, eq_iff_exists := λ x y, ⟨λ h, ⟨1, by simpa using algebra_map_injective C A L h⟩, λ ⟨c, hc⟩, - congr_arg (algebra_map _ L) (mul_right_cancel₀ (mem_non_zero_divisors_iff_ne_zero.mp c.2) hc)⟩ } + congr_arg (algebra_map _ L) (mul_left_cancel₀ (mem_non_zero_divisors_iff_ne_zero.mp c.2) hc)⟩ } variables (K L) @@ -420,7 +420,7 @@ begin have mk_yz_eq : is_localization.mk' L y' z' = is_localization.mk' L y ⟨_, hz0'⟩, { rw [algebra.smul_def, mul_comm _ y, mul_comm _ y', ← set_like.coe_mk (algebra_map R S z) hz0'] at yz_eq, - exact is_localization.mk'_eq_of_eq yz_eq.symm }, + exact is_localization.mk'_eq_of_eq (by rw [mul_comm _ y, mul_comm _ y', yz_eq]), }, suffices hy : algebra_map S L (a * y) ∈ submodule.span K (⇑(algebra_map S L) '' b), { rw [mk_yz_eq, is_fraction_ring.mk'_eq_div, set_like.coe_mk, ← is_scalar_tower.algebra_map_apply, is_scalar_tower.algebra_map_apply R K L, diff --git a/src/ring_theory/localization/localization_localization.lean b/src/ring_theory/localization/localization_localization.lean index a21991101fca7..86ed623a3cb6c 100644 --- a/src/ring_theory/localization/localization_localization.lean +++ b/src/ring_theory/localization/localization_localization.lean @@ -87,7 +87,7 @@ end lemma localization_localization_eq_iff_exists [is_localization N T] (x y : R) : algebra_map R T x = algebra_map R T y ↔ - ∃ (c : localization_localization_submodule M N), x * c = y * c := + ∃ (c : localization_localization_submodule M N), ↑c * x = ↑c * y := begin rw [is_scalar_tower.algebra_map_apply R S T, is_scalar_tower.algebra_map_apply R S T, is_localization.eq_iff_exists N T], @@ -95,19 +95,20 @@ begin { rintros ⟨z, eq₁⟩, rcases is_localization.surj M (z : S) with ⟨⟨z', s⟩, eq₂⟩, dsimp only at eq₂, - obtain ⟨c, eq₃ : x * z' * ↑ c = y * z' * ↑ c⟩ := (is_localization.eq_iff_exists M S).mp _, - swap, { rw [ring_hom.map_mul, ring_hom.map_mul, ← eq₂, ← mul_assoc, ← mul_assoc, ← eq₁] }, - use z' * c, + obtain ⟨c, eq₃ : ↑c * (x * z') = ↑c * (y * z')⟩ := (is_localization.eq_iff_exists M S).mp _, + swap, + { rw [map_mul, map_mul, ←eq₂, ←mul_assoc, ←mul_assoc, mul_comm _ ↑z, eq₁, mul_comm _ ↑z] }, + use c * z', { rw mem_localization_localization_submodule, - refine ⟨z, s * c, _⟩, - rw [ring_hom.map_mul, ← eq₂, mul_assoc, ← ring_hom.map_mul, submonoid.coe_mul] }, - { simpa only [mul_assoc] using eq₃ } }, - { rintro ⟨⟨c, hc⟩, eq₁ : x * c = y * c⟩, + refine ⟨z, c * s, _⟩, + rw [map_mul, ← eq₂, submonoid.coe_mul, map_mul, mul_left_comm] }, + { rwa [mul_comm _ z', mul_comm _ z', ←mul_assoc, ←mul_assoc] at eq₃ } }, + { rintro ⟨⟨c, hc⟩, eq₁ : c * x = c * y⟩, rw mem_localization_localization_submodule at hc, rcases hc with ⟨z₁, z, eq₂⟩, use z₁, - refine (is_localization.map_units S z).mul_left_inj.mp _, - rw [mul_assoc, mul_assoc, ← eq₂, ← ring_hom.map_mul, ← ring_hom.map_mul, eq₁] } + refine (is_localization.map_units S z).mul_right_inj.mp _, + rw [←mul_assoc, mul_comm _ ↑z₁, ←eq₂, ←map_mul, eq₁, map_mul, eq₂, ←mul_assoc, mul_comm _ ↑z₁] } end /-- @@ -217,7 +218,7 @@ lemma is_localization_of_submonoid_le eq_iff_exists := λ x₁ x₂, begin obtain ⟨⟨y₁, s₁⟩, e₁⟩ := is_localization.surj M x₁, obtain ⟨⟨y₂, s₂⟩, e₂⟩ := is_localization.surj M x₂, - refine iff.trans _ (set.exists_image_iff (algebra_map R S) N (λ c, x₁ * c = x₂ * c)).symm, + refine iff.trans _ (set.exists_image_iff (algebra_map R S) N (λ c, c * x₁ = c * x₂)).symm, dsimp only at e₁ e₂ ⊢, suffices : algebra_map R T (y₁ * s₂) = algebra_map R T (y₂ * s₁) ↔ ∃ (a : N), algebra_map R S (a * (y₁ * s₂)) = algebra_map R S (a * (y₂ * s₁)), @@ -254,7 +255,7 @@ lemma is_localization_of_is_exists_mul_mem (M N : submonoid R) [is_localization rintros ⟨x, h⟩, obtain ⟨m, hm⟩ := h' x, refine ⟨⟨_, hm⟩, _⟩, - simp [mul_comm m, ← mul_assoc, h] + simp [h, mul_assoc], end } end localization_localization @@ -300,7 +301,7 @@ begin intro hx', apply @zero_ne_one S, rw [← (algebra_map R S).map_one, ← @mk'_one R _ M, @comm _ eq, mk'_eq_zero_iff], - exact ⟨⟨_, hx⟩, (one_mul x).symm ▸ hx'⟩, + exact ⟨⟨x, hx⟩, by simp [hx'] ⟩, end end is_fraction_ring diff --git a/src/ring_theory/localization/num_denom.lean b/src/ring_theory/localization/num_denom.lean index db61a20f6d821..2a33e89e230b2 100644 --- a/src/ring_theory/localization/num_denom.lean +++ b/src/ring_theory/localization/num_denom.lean @@ -74,7 +74,7 @@ lemma num_mul_denom_eq_num_iff_eq' {x y : K} : lemma num_mul_denom_eq_num_mul_denom_iff_eq {x y : K} : num A y * denom A x = num A x * denom A y ↔ x = y := -⟨λ h, by simpa only [mk'_num_denom] using mk'_eq_of_eq h, +⟨λ h, by simpa only [mk'_num_denom] using mk'_eq_of_eq' h, λ h, by rw h⟩ lemma eq_zero_of_num_eq_zero {x : K} (h : num A x = 0) : x = 0 := diff --git a/src/ring_theory/ore_localization/basic.lean b/src/ring_theory/ore_localization/basic.lean index 0a968d6ae647e..86a19adfaea4a 100644 --- a/src/ring_theory/ore_localization/basic.lean +++ b/src/ring_theory/ore_localization/basic.lean @@ -409,10 +409,15 @@ protected def localization_map : S.localization_map R[S⁻¹] := { intro h, rw [numerator_hom_apply, numerator_hom_apply, ore_div_eq_iff] at h, rcases h with ⟨u, v, h₁, h₂⟩, dsimp at h₂, - rw [one_mul, one_mul] at h₂, subst h₂, use u, exact h₁.symm }, + rw [one_mul, one_mul] at h₂, + subst h₂, + use u, + simpa only [mul_comm] using h₁.symm }, { rintro ⟨s, h⟩, rw [numerator_hom_apply, numerator_hom_apply, ore_div_eq_iff], - use s, use s, simp [h, one_mul] } + refine ⟨s, s, _, _⟩, + { simpa [mul_comm] using h.symm }, + { simp [one_mul]} } end } /-- If `R` is commutative, Ore localization and monoid localization are isomorphic. -/ diff --git a/src/ring_theory/ring_hom/surjective.lean b/src/ring_theory/ring_hom/surjective.lean index a2687a5774905..ae8a364937991 100644 --- a/src/ring_theory/ring_hom/surjective.lean +++ b/src/ring_theory/ring_hom/surjective.lean @@ -69,10 +69,10 @@ begin obtain ⟨z, ⟨_, n, rfl⟩, rfl⟩ := is_localization.mk'_surjective (submonoid.powers (r : R)) y, erw [is_localization.map_mk', is_localization.eq] at hy, obtain ⟨⟨_, m, rfl⟩, hm⟩ := hy, - dsimp at hm, - simp_rw [_root_.mul_assoc, _root_.one_mul, ← map_pow, ← f.map_mul, ← pow_add, mul_comm x] at hm, - rw map_pow at hm, - refine ⟨n + m, _, hm⟩ } + refine ⟨m + n, _⟩, + dsimp at hm ⊢, + simp_rw [_root_.one_mul, ← _root_.mul_assoc, ← map_pow, ← f.map_mul, ← pow_add, map_pow] at hm, + exact ⟨_, hm⟩ } end end ring_hom diff --git a/src/ring_theory/valuation/valuation_subring.lean b/src/ring_theory/valuation/valuation_subring.lean index c66036bcc8089..49e106e62a016 100644 --- a/src/ring_theory/valuation/valuation_subring.lean +++ b/src/ring_theory/valuation/valuation_subring.lean @@ -105,7 +105,7 @@ instance : is_fraction_ring A K := exact mem_non_zero_divisors_iff_ne_zero.2 (λ c, h (inv_eq_zero.mp (congr_arg coe c))) }, end, eq_iff_exists := λ a b, ⟨ λ h, ⟨1, by { ext, simpa using h }⟩, λ ⟨c, h⟩, - congr_arg coe ((mul_eq_mul_right_iff.1 h).resolve_right (non_zero_divisors.ne_zero c.2)) ⟩ } + congr_arg coe ((mul_eq_mul_left_iff.1 h).resolve_right (non_zero_divisors.ne_zero c.2)) ⟩ } /-- The value group of the valuation associated to `A`. Note: it is actually a group with zero. -/ @[derive linear_ordered_comm_group_with_zero] diff --git a/src/set_theory/surreal/dyadic.lean b/src/set_theory/surreal/dyadic.lean index 74b9f7128fc4c..fa55c3c1c7a2d 100644 --- a/src/set_theory/surreal/dyadic.lean +++ b/src/set_theory/surreal/dyadic.lean @@ -185,7 +185,7 @@ def dyadic_map : localization.away (2 : ℤ) →+ surreal := begin intros m₁ m₂ n₁ n₂ h₁, obtain ⟨⟨n₃, y₃, hn₃⟩, h₂⟩ := localization.r_iff_exists.mp h₁, - simp only [subtype.coe_mk, mul_eq_mul_right_iff] at h₂, + simp only [subtype.coe_mk, mul_eq_mul_left_iff] at h₂, cases h₂, { simp only, obtain ⟨a₁, ha₁⟩ := n₁.prop, @@ -195,7 +195,7 @@ def dyadic_map : localization.away (2 : ℤ) →+ surreal := have h₂ : 1 < (2 : ℤ).nat_abs, from one_lt_two, rw [hn₁, hn₂, submonoid.log_pow_int_eq_self h₂, submonoid.log_pow_int_eq_self h₂], apply dyadic_aux, - rwa [ha₁, ha₂] }, + rwa [ha₁, ha₂, mul_comm, mul_comm m₂] }, { have : (1 : ℤ) ≤ 2 ^ y₃ := by exact_mod_cast nat.one_le_pow y₃ 2 nat.succ_pos', linarith } end, From 800d3d45716ea9f4a8c67ffc8b53da951593141e Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Sun, 22 Jan 2023 11:32:06 +0000 Subject: [PATCH 0319/1291] feat(set_theory/zfc/basic): better docs on `univ_not_mem_univ` (#18234) This theorem is somewhat odd, but phrases a very important result: there is no ZFC universal set. The new docs better clarify this. --- src/set_theory/zfc/basic.lean | 5 ++++- 1 file changed, 4 insertions(+), 1 deletion(-) diff --git a/src/set_theory/zfc/basic.lean b/src/set_theory/zfc/basic.lean index 2c38bbbcc4d59..975ff49883ea7 100644 --- a/src/set_theory/zfc/basic.lean +++ b/src/set_theory/zfc/basic.lean @@ -894,7 +894,10 @@ instance : is_asymm Class (∈) := mem_wf.is_asymm theorem mem_asymm {x y : Class} : x ∈ y → y ∉ x := asymm theorem mem_irrefl (x : Class) : x ∉ x := irrefl x -/-- There is no universal set. -/ +/-- **There is no universal set.** + +This is stated as `univ ∉ univ`, meaning that `univ` (the class of all sets) is proper (does not +belong to the class of all sets). -/ theorem univ_not_mem_univ : univ ∉ univ := mem_irrefl _ /-- Convert a conglomerate (a collection of classes) into a class -/ From ae5357f97a127441d38a62ff10862a40059a7f13 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sun, 22 Jan 2023 11:32:07 +0000 Subject: [PATCH 0320/1291] refactor(data/finsupp/basic): split out alist results (#18250) This file is getting quite long, and nothing else builds upon these results. These lemmas and definitions are copied without modification. They were originally from #15443. --- src/data/finsupp/alist.lean | 107 ++++++++++++++++++++++++++++++++++++ src/data/finsupp/basic.lean | 94 ------------------------------- 2 files changed, 107 insertions(+), 94 deletions(-) create mode 100644 src/data/finsupp/alist.lean diff --git a/src/data/finsupp/alist.lean b/src/data/finsupp/alist.lean new file mode 100644 index 0000000000000..48540fb78e393 --- /dev/null +++ b/src/data/finsupp/alist.lean @@ -0,0 +1,107 @@ +/- +Copyright (c) 2022 Violeta Hernández. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Violeta Hernández +-/ +import data.finsupp.basic +import data.list.alist + +/-! +# Connections between `finsupp` and `alist` + +## Main definitions + +* `finsupp.to_alist` +* `alist.lookup_finsupp`: converts an association list into a finitely supported function + via `alist.lookup`, sending absent keys to zero. + +-/ + +namespace finsupp +variables {α M : Type*} [has_zero M] + +/-- Produce an association list for the finsupp over its support using choice. -/ +@[simps] noncomputable def to_alist (f : α →₀ M) : alist (λ x : α, M) := +⟨f.graph.to_list.map prod.to_sigma, begin + rw [list.nodupkeys, list.keys, list.map_map, prod.fst_comp_to_sigma, list.nodup_map_iff_inj_on], + { rintros ⟨b, m⟩ hb ⟨c, n⟩ hc (rfl : b = c), + rw [finset.mem_to_list, finsupp.mem_graph_iff] at hb hc, + dsimp at hb hc, + rw [←hc.1, hb.1] }, + { apply finset.nodup_to_list } +end⟩ + +@[simp] lemma to_alist_keys_to_finset [decidable_eq α] (f : α →₀ M) : + f.to_alist.keys.to_finset = f.support := +by { ext, simp [to_alist, alist.mem_keys, alist.keys, list.keys] } + +@[simp] lemma mem_to_alist {f : α →₀ M} {x : α} : x ∈ f.to_alist ↔ f x ≠ 0 := +begin + classical, + rw [alist.mem_keys, ←list.mem_to_finset, to_alist_keys_to_finset, mem_support_iff] +end + +end finsupp + +namespace alist +variables {α M : Type*} [has_zero M] +open list + +/-- Converts an association list into a finitely supported function via `alist.lookup`, sending +absent keys to zero. -/ +noncomputable def lookup_finsupp (l : alist (λ x : α, M)) : α →₀ M := +{ support := by haveI := classical.dec_eq α; haveI := classical.dec_eq M; exact + (l.1.filter $ λ x, sigma.snd x ≠ 0).keys.to_finset, + to_fun := λ a, by haveI := classical.dec_eq α; exact (l.lookup a).get_or_else 0, + mem_support_to_fun := λ a, begin + classical, + simp_rw [mem_to_finset, list.mem_keys, list.mem_filter, ←mem_lookup_iff], + cases lookup a l; + simp + end } + +@[simp] lemma lookup_finsupp_apply [decidable_eq α] (l : alist (λ x : α, M)) (a : α) : + l.lookup_finsupp a = (l.lookup a).get_or_else 0 := +by convert rfl + +@[simp] lemma lookup_finsupp_support [decidable_eq α] [decidable_eq M] (l : alist (λ x : α, M)) : + l.lookup_finsupp.support = (l.1.filter $ λ x, sigma.snd x ≠ 0).keys.to_finset := +by convert rfl + +lemma lookup_finsupp_eq_iff_of_ne_zero [decidable_eq α] + {l : alist (λ x : α, M)} {a : α} {x : M} (hx : x ≠ 0) : + l.lookup_finsupp a = x ↔ x ∈ l.lookup a := +by { rw lookup_finsupp_apply, cases lookup a l with m; simp [hx.symm] } + +lemma lookup_finsupp_eq_zero_iff [decidable_eq α] {l : alist (λ x : α, M)} {a : α} : + l.lookup_finsupp a = 0 ↔ a ∉ l ∨ (0 : M) ∈ l.lookup a := +by { rw [lookup_finsupp_apply, ←lookup_eq_none], cases lookup a l with m; simp } + +@[simp] lemma empty_lookup_finsupp : lookup_finsupp (∅ : alist (λ x : α, M)) = 0 := +by { classical, ext, simp } + +@[simp] lemma insert_lookup_finsupp [decidable_eq α] (l : alist (λ x : α, M)) (a : α) (m : M) : + (l.insert a m).lookup_finsupp = l.lookup_finsupp.update a m := +by { ext b, by_cases h : b = a; simp [h] } + +@[simp] lemma singleton_lookup_finsupp (a : α) (m : M) : + (singleton a m).lookup_finsupp = finsupp.single a m := +by { classical, simp [←alist.insert_empty] } + +@[simp] lemma _root_.finsupp.to_alist_lookup_finsupp (f : α →₀ M) : f.to_alist.lookup_finsupp = f := +begin + ext, + by_cases h : f a = 0, + { suffices : f.to_alist.lookup a = none, + { simp [h, this] }, + { simp [lookup_eq_none, h] } }, + { suffices : f.to_alist.lookup a = some (f a), + { simp [h, this] }, + { apply mem_lookup_iff.2, + simpa using h } } +end + +lemma lookup_finsupp_surjective : function.surjective (@lookup_finsupp α M _) := +λ f, ⟨_, finsupp.to_alist_lookup_finsupp f⟩ + +end alist diff --git a/src/data/finsupp/basic.lean b/src/data/finsupp/basic.lean index ad52601377153..c3ea2d6e43d42 100644 --- a/src/data/finsupp/basic.lean +++ b/src/data/finsupp/basic.lean @@ -7,7 +7,6 @@ import algebra.big_operators.finsupp import algebra.hom.group_action import algebra.regular.smul import data.finset.preimage -import data.list.alist import data.rat.big_operators /-! @@ -16,8 +15,6 @@ import data.rat.big_operators ## Main declarations * `finsupp.graph`: the finset of input and output pairs with non-zero outputs. -* `alist.lookup_finsupp`: converts an association list into a finitely supported function - via `alist.lookup`, sending absent keys to zero. * `finsupp.map_range.equiv`: `finsupp.map_range` as an equiv. * `finsupp.map_domain`: maps the domain of a `finsupp` by a function and by summing. * `finsupp.comap_domain`: postcomposition of a `finsupp` with a function injective on the preimage @@ -107,101 +104,10 @@ end @[simp] lemma graph_eq_empty {f : α →₀ M} : f.graph = ∅ ↔ f = 0 := (graph_injective α M).eq_iff' graph_zero -/-- Produce an association list for the finsupp over its support using choice. -/ -@[simps] def to_alist (f : α →₀ M) : alist (λ x : α, M) := -⟨f.graph.to_list.map prod.to_sigma, begin - rw [list.nodupkeys, list.keys, list.map_map, prod.fst_comp_to_sigma, list.nodup_map_iff_inj_on], - { rintros ⟨b, m⟩ hb ⟨c, n⟩ hc (rfl : b = c), - rw [mem_to_list, finsupp.mem_graph_iff] at hb hc, - dsimp at hb hc, - rw [←hc.1, hb.1] }, - { apply nodup_to_list } -end⟩ - -@[simp] lemma to_alist_keys_to_finset [decidable_eq α] (f : α →₀ M) : - f.to_alist.keys.to_finset = f.support := -by { ext, simp [to_alist, alist.mem_keys, alist.keys, list.keys] } - -@[simp] lemma mem_to_alist {f : α →₀ M} {x : α} : x ∈ f.to_alist ↔ f x ≠ 0 := -begin - classical, - rw [alist.mem_keys, ←list.mem_to_finset, to_alist_keys_to_finset, mem_support_iff] -end - end graph end finsupp -/-! ### Declarations about `alist.lookup_finsupp` -/ - -section lookup_finsupp - -variable [has_zero M] - -namespace alist -open list - -/-- Converts an association list into a finitely supported function via `alist.lookup`, sending -absent keys to zero. -/ -def lookup_finsupp (l : alist (λ x : α, M)) : α →₀ M := -{ support := by haveI := classical.dec_eq α; haveI := classical.dec_eq M; exact - (l.1.filter $ λ x, sigma.snd x ≠ 0).keys.to_finset, - to_fun := λ a, by haveI := classical.dec_eq α; exact (l.lookup a).get_or_else 0, - mem_support_to_fun := λ a, begin - classical, - simp_rw [mem_to_finset, list.mem_keys, list.mem_filter, ←mem_lookup_iff], - cases lookup a l; - simp - end } - -@[simp] lemma lookup_finsupp_apply [decidable_eq α] (l : alist (λ x : α, M)) (a : α) : - l.lookup_finsupp a = (l.lookup a).get_or_else 0 := -by convert rfl - -@[simp] lemma lookup_finsupp_support [decidable_eq α] [decidable_eq M] (l : alist (λ x : α, M)) : - l.lookup_finsupp.support = (l.1.filter $ λ x, sigma.snd x ≠ 0).keys.to_finset := -by convert rfl - -lemma lookup_finsupp_eq_iff_of_ne_zero [decidable_eq α] - {l : alist (λ x : α, M)} {a : α} {x : M} (hx : x ≠ 0) : - l.lookup_finsupp a = x ↔ x ∈ l.lookup a := -by { rw lookup_finsupp_apply, cases lookup a l with m; simp [hx.symm] } - -lemma lookup_finsupp_eq_zero_iff [decidable_eq α] {l : alist (λ x : α, M)} {a : α} : - l.lookup_finsupp a = 0 ↔ a ∉ l ∨ (0 : M) ∈ l.lookup a := -by { rw [lookup_finsupp_apply, ←lookup_eq_none], cases lookup a l with m; simp } - -@[simp] lemma empty_lookup_finsupp : lookup_finsupp (∅ : alist (λ x : α, M)) = 0 := -by { classical, ext, simp } - -@[simp] lemma insert_lookup_finsupp [decidable_eq α] (l : alist (λ x : α, M)) (a : α) (m : M) : - (l.insert a m).lookup_finsupp = l.lookup_finsupp.update a m := -by { ext b, by_cases h : b = a; simp [h] } - -@[simp] lemma singleton_lookup_finsupp (a : α) (m : M) : - (singleton a m).lookup_finsupp = finsupp.single a m := -by { classical, simp [←alist.insert_empty] } - -@[simp] lemma _root_.finsupp.to_alist_lookup_finsupp (f : α →₀ M) : f.to_alist.lookup_finsupp = f := -begin - ext, - by_cases h : f a = 0, - { suffices : f.to_alist.lookup a = none, - { simp [h, this] }, - { simp [lookup_eq_none, h] } }, - { suffices : f.to_alist.lookup a = some (f a), - { simp [h, this] }, - { apply mem_lookup_iff.2, - simpa using h } } -end - -lemma lookup_finsupp_surjective : surjective (@lookup_finsupp α M _) := -λ f, ⟨_, finsupp.to_alist_lookup_finsupp f⟩ - -end alist - -end lookup_finsupp - /-! ### Declarations about `map_range` -/ section map_range From fc3dc609af89ac2ca2a501541e0647b6c179a0f9 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sun, 22 Jan 2023 11:32:08 +0000 Subject: [PATCH 0321/1291] chore(order/filter/basic): fix 2 typos (#18253) --- src/order/filter/basic.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/order/filter/basic.lean b/src/order/filter/basic.lean index ca2258da2cbd6..231daf4797bbf 100644 --- a/src/order/filter/basic.lean +++ b/src/order/filter/basic.lean @@ -1765,7 +1765,7 @@ lemma _root_.function.semiconj.filter_map {f : α → β} {ga : α → α} {gb : (h : function.semiconj f ga gb) : function.semiconj (map f) (map ga) (map gb) := map_comm h.comp_eq -lemma _root_.commute.filter_map {f g : α → α} (h : function.commute f g) : +lemma _root_.function.commute.filter_map {f g : α → α} (h : function.commute f g) : function.commute (map f) (map g) := h.filter_map @@ -1773,7 +1773,7 @@ lemma _root_.function.semiconj.filter_comap {f : α → β} {ga : α → α} {gb (h : function.semiconj f ga gb) : function.semiconj (comap f) (comap gb) (comap ga) := comap_comm h.comp_eq.symm -lemma _root_.commute.filter_comap {f g : α → α} (h : function.commute f g) : +lemma _root_.function.commute.filter_comap {f g : α → α} (h : function.commute f g) : function.commute (comap f) (comap g) := h.filter_comap From 71efed1bb08264afaee50239e168e0199d29577d Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sun, 22 Jan 2023 11:32:09 +0000 Subject: [PATCH 0322/1291] fix(order/filter/basic): fix a diamond (#18255) Fix non-defeq `default` in `filter.inhabited` and `filter.unique`. --- src/order/filter/basic.lean | 7 +++---- 1 file changed, 3 insertions(+), 4 deletions(-) diff --git a/src/order/filter/basic.lean b/src/order/filter/basic.lean index 231daf4797bbf..79f73a8030deb 100644 --- a/src/order/filter/basic.lean +++ b/src/order/filter/basic.lean @@ -244,9 +244,6 @@ def principal (s : set α) : filter α := localized "notation (name := filter.principal) `𝓟` := filter.principal" in filter -instance : inhabited (filter α) := -⟨𝓟 ∅⟩ - @[simp] lemma mem_principal {s t : set α} : s ∈ 𝓟 t ↔ t ⊆ s := iff.rfl lemma mem_principal_self (s : set α) : s ∈ 𝓟 s := subset.rfl @@ -429,6 +426,8 @@ instance : complete_lattice (filter α) := original_complete_lattice.copy (set.ext_iff.1 (sInter_image _ _) x).symm}) /- Inf -/ _ rfl +instance : inhabited (filter α) := ⟨⊥⟩ + end complete_lattice /-- A filter is `ne_bot` if it is not equal to `⊥`, or equivalently the empty set @@ -660,7 +659,7 @@ end /-- There is exactly one filter on an empty type. -/ instance unique [is_empty α] : unique (filter α) := -{ default := ⊥, uniq := filter_eq_bot_of_is_empty } +{ to_inhabited := filter.inhabited, uniq := filter_eq_bot_of_is_empty } /-- There are only two filters on a `subsingleton`: `⊥` and `⊤`. If the type is empty, then they are equal. -/ From 386c02e7c94317d92b78c6ff98462614bcbb3699 Mon Sep 17 00:00:00 2001 From: Mark Andrew Gerads Date: Sun, 22 Jan 2023 14:11:20 +0000 Subject: [PATCH 0323/1291] feat(data/nat/hyperoperation): Defined hyperoperation and added related lemmas (#18116) Added the hyperoperation sequence as defined here: https://en.wikipedia.org/wiki/Hyperoperation Proved main lemmas. --- src/data/nat/hyperoperation.lean | 121 +++++++++++++++++++++++++++++++ 1 file changed, 121 insertions(+) create mode 100644 src/data/nat/hyperoperation.lean diff --git a/src/data/nat/hyperoperation.lean b/src/data/nat/hyperoperation.lean new file mode 100644 index 0000000000000..ca0f6f25a3b5e --- /dev/null +++ b/src/data/nat/hyperoperation.lean @@ -0,0 +1,121 @@ +/- +Copyright (c) 2023 Mark Andrew Gerads. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Mark Andrew Gerads, Junyan Xu, Eric Wieser +-/ +import tactic.ring +import data.nat.parity + +/-! +# Hyperoperation sequence + +This file defines the Hyperoperation sequence. +`hyperoperation 0 m k = k + 1` +`hyperoperation 1 m k = m + k` +`hyperoperation 2 m k = m * k` +`hyperoperation 3 m k = m ^ k` +`hyperoperation (n + 3) m 0 = 1` +`hyperoperation (n + 1) m (k + 1) = hyperoperation n m (hyperoperation (n + 1) m k)` + +## References + +* + +## Tags + +hyperoperation +-/ + +/-- +Implementation of the hyperoperation sequence +where `hyperoperation n m k` is the `n`th hyperoperation between `m` and `k`. +-/ +def hyperoperation : ℕ → ℕ → ℕ → ℕ +| 0 _ k := k + 1 +| 1 m 0 := m +| 2 _ 0 := 0 +| (n + 3) _ 0 := 1 +| (n + 1) m (k + 1) := hyperoperation n m (hyperoperation (n + 1) m k) + +-- Basic hyperoperation lemmas + +@[simp] lemma hyperoperation_zero (m : ℕ) : hyperoperation 0 m = nat.succ := +funext $ λ k, by rw [hyperoperation, nat.succ_eq_add_one] + +lemma hyperoperation_ge_three_eq_one (n m : ℕ) : hyperoperation (n + 3) m 0 = 1 := +by rw hyperoperation + +lemma hyperoperation_recursion (n m k : ℕ) : + hyperoperation (n + 1) m (k + 1) = hyperoperation n m (hyperoperation (n + 1) m k) := +by obtain (_|_|_) := n; rw hyperoperation + +-- Interesting hyperoperation lemmas + +@[simp] lemma hyperoperation_one : hyperoperation 1 = (+) := +begin + ext m k, + induction k with bn bih, + { rw [nat_add_zero m, hyperoperation], }, + { rw [hyperoperation_recursion, bih, hyperoperation_zero], + exact nat.add_assoc m bn 1, }, +end + +@[simp] lemma hyperoperation_two : hyperoperation 2 = (*) := +begin + ext m k, + induction k with bn bih, + { rw hyperoperation, + exact (nat.mul_zero m).symm, }, + { rw [hyperoperation_recursion, hyperoperation_one, bih], + ring, }, +end + +@[simp] lemma hyperoperation_three : hyperoperation 3 = (^) := +begin + ext m k, + induction k with bn bih, + { rw hyperoperation_ge_three_eq_one, + exact (pow_zero m).symm, }, + { rw [hyperoperation_recursion, hyperoperation_two, bih], + exact (pow_succ m bn).symm, }, +end + +lemma hyperoperation_ge_two_eq_self (n m : ℕ) : hyperoperation (n + 2) m 1 = m := +begin + induction n with nn nih, + { rw hyperoperation_two, + ring, }, + { rw [hyperoperation_recursion, hyperoperation_ge_three_eq_one, nih], }, +end + +lemma hyperoperation_two_two_eq_four (n : ℕ) : hyperoperation (n + 1) 2 2 = 4 := +begin + induction n with nn nih, + { rw hyperoperation_one, }, + { rw [hyperoperation_recursion, hyperoperation_ge_two_eq_self, nih], }, +end + +lemma hyperoperation_ge_three_one (n : ℕ) : ∀ (k : ℕ), hyperoperation (n + 3) 1 k = 1 := +begin + induction n with nn nih, + { intros k, + rw [hyperoperation_three, one_pow], }, + { intros k, + cases k, + { rw hyperoperation_ge_three_eq_one, }, + { rw [hyperoperation_recursion, nih], }, }, +end + +lemma hyperoperation_ge_four_zero (n k : ℕ) : + hyperoperation (n + 4) 0 k = if (even k) then 1 else 0 := +begin + induction k with kk kih, + { rw hyperoperation_ge_three_eq_one, + simp only [even_zero, if_true], }, + { rw hyperoperation_recursion, + rw kih, + simp_rw nat.even_add_one, + split_ifs, + { exact hyperoperation_ge_two_eq_self (n + 1) 0, }, + { exact hyperoperation_ge_three_eq_one n 0, }, }, +end From 9df54c1b9dd5b6d6ed1472bfb6c10819aa95b5d7 Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Sun, 22 Jan 2023 14:11:21 +0000 Subject: [PATCH 0324/1291] feat(special_functions/trigonometric/bounds): Add the inequality `cos x < 1 / sqrt (x ^ 2 + 1)` (#18189) --- .../trigonometric/bounds.lean | 47 ++++++++++++++++++- 1 file changed, 45 insertions(+), 2 deletions(-) diff --git a/src/analysis/special_functions/trigonometric/bounds.lean b/src/analysis/special_functions/trigonometric/bounds.lean index d0ad2627ea9e4..e185a6bbf1efe 100644 --- a/src/analysis/special_functions/trigonometric/bounds.lean +++ b/src/analysis/special_functions/trigonometric/bounds.lean @@ -19,6 +19,9 @@ Here we prove the following: * `sin_lt`: for `x > 0` we have `sin x < x`. * `sin_gt_sub_cube`: For `0 < x ≤ 1` we have `x - x ^ 3 / 4 < sin x`. * `lt_tan`: for `0 < x < π/2` we have `x < tan x`. +* `cos_le_one_div_sqrt_sq_add_one` and `cos_lt_one_div_sqrt_sq_add_one`: for + `-3 * π / 2 ≤ x ≤ 3 * π / 2`, we have `cos x ≤ 1 / sqrt (x ^ 2 + 1)`, with strict inequality if + `x ≠ 0`. (This bound is not quite optimal, but not far off) ## Tags @@ -70,11 +73,11 @@ lemma deriv_tan_sub_id (x : ℝ) (h : cos x ≠ 0) : deriv (λ y : ℝ, tan y - y) x = 1 / cos x ^ 2 - 1 := has_deriv_at.deriv $ by simpa using (has_deriv_at_tan h).add (has_deriv_at_id x).neg -/-- For all `0 ≤ x < π/2` we have `x < tan x`. +/-- For all `0 < x < π/2` we have `x < tan x`. This is proved by checking that the function `tan x - x` vanishes at zero and has non-negative derivative. -/ -theorem lt_tan (x : ℝ) (h1 : 0 < x) (h2 : x < π / 2) : x < tan x := +theorem lt_tan {x : ℝ} (h1 : 0 < x) (h2 : x < π / 2) : x < tan x := begin let U := Ico 0 (π / 2), @@ -120,4 +123,44 @@ begin simpa only [tan_zero, sub_zero, sub_pos] using mono zero_in_U x_in_U h1 end +lemma le_tan {x : ℝ} (h1 : 0 ≤ x) (h2 : x < π / 2) : x ≤ tan x := +begin + rcases eq_or_lt_of_le h1 with rfl | h1', + { rw tan_zero }, + { exact le_of_lt (lt_tan h1' h2) } +end + +lemma cos_lt_one_div_sqrt_sq_add_one {x : ℝ} + (hx1 : -(3 * π / 2) ≤ x) (hx2 : x ≤ 3 * π / 2) (hx3 : x ≠ 0) : + cos x < 1 / sqrt (x ^ 2 + 1) := +begin + suffices : ∀ {y : ℝ} (hy1 : 0 < y) (hy2 : y ≤ 3 * π / 2), cos y < 1 / sqrt (y ^ 2 + 1), + { rcases lt_or_lt_iff_ne.mpr hx3.symm, + { exact this h hx2 }, + { convert this (by linarith : 0 < -x) (by linarith) using 1, + { rw cos_neg }, { rw neg_sq } } }, + intros y hy1 hy2, + have hy3 : 0 < y ^ 2 + 1, by linarith [sq_nonneg y], + rcases lt_or_le y (π / 2) with hy2' | hy1', + { -- Main case : `0 < y < π / 2` + have hy4 : 0 < cos y := cos_pos_of_mem_Ioo ⟨by linarith, hy2'⟩, + rw [←abs_of_nonneg (cos_nonneg_of_mem_Icc ⟨by linarith, hy2'.le⟩), + ←abs_of_nonneg (one_div_nonneg.mpr (sqrt_nonneg _)), ←sq_lt_sq, div_pow, one_pow, + sq_sqrt hy3.le, lt_one_div (pow_pos hy4 _) hy3, ←inv_one_add_tan_sq hy4.ne', one_div, inv_inv, + add_comm, add_lt_add_iff_left, sq_lt_sq, abs_of_pos hy1, + abs_of_nonneg (tan_nonneg_of_nonneg_of_le_pi_div_two hy1.le hy2'.le)], + exact real.lt_tan hy1 hy2' }, + { -- Easy case : `π / 2 ≤ y ≤ 3 * π / 2` + refine lt_of_le_of_lt _ (one_div_pos.mpr $ sqrt_pos_of_pos hy3), + exact cos_nonpos_of_pi_div_two_le_of_le hy1' (by linarith [pi_pos]) } +end + +lemma cos_le_one_div_sqrt_sq_add_one {x : ℝ} (hx1 : -(3 * π / 2) ≤ x) (hx2 : x ≤ 3 * π / 2) : + cos x ≤ 1 / sqrt (x ^ 2 + 1) := +begin + rcases eq_or_ne x 0 with rfl | hx3, + { simp }, + { exact (cos_lt_one_div_sqrt_sq_add_one hx1 hx2 hx3).le } +end + end real From 160ef3e338a2a4f21a280e4c152d4016156e516d Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sun, 22 Jan 2023 14:11:23 +0000 Subject: [PATCH 0325/1291] chore(analysis/von_neumann_algebra/basic): add missing API lemmas (#18227) There's no interesting math here, just some missing bundling --- src/algebra/star/subalgebra.lean | 7 +++- src/analysis/von_neumann_algebra/basic.lean | 43 +++++++++++++++++++-- 2 files changed, 46 insertions(+), 4 deletions(-) diff --git a/src/algebra/star/subalgebra.lean b/src/algebra/star/subalgebra.lean index 8e33aa409a99a..b2c0793876a2f 100644 --- a/src/algebra/star/subalgebra.lean +++ b/src/algebra/star/subalgebra.lean @@ -226,10 +226,15 @@ end map section centralizer variables (R) {A} +lemma _root_.set.star_mem_centralizer {a : A} {s : set A} + (h : ∀ (a : A), a ∈ s → star a ∈ s) (ha : a ∈ set.centralizer s) : + star a ∈ set.centralizer s := +λ y hy, by simpa using congr_arg star (ha _ (h _ hy)).symm + /-- The centralizer, or commutant, of a *-closed set as star subalgebra. -/ def centralizer (s : set A) (w : ∀ (a : A), a ∈ s → star a ∈ s) : star_subalgebra R A := -{ star_mem' := λ x xm y hy, by simpa using congr_arg star (xm _ (w _ hy)).symm, +{ star_mem' := λ x, set.star_mem_centralizer w, ..subalgebra.centralizer R s, } @[simp] diff --git a/src/analysis/von_neumann_algebra/basic.lean b/src/analysis/von_neumann_algebra/basic.lean index e6c427a9f1301..365c0744733d1 100644 --- a/src/analysis/von_neumann_algebra/basic.lean +++ b/src/analysis/von_neumann_algebra/basic.lean @@ -67,7 +67,8 @@ and instead will use `⊤ : von_neumann_algebra H`. @[nolint has_nonempty_instance] structure von_neumann_algebra (H : Type u) [inner_product_space ℂ H] [complete_space H] extends star_subalgebra ℂ (H →L[ℂ] H) := -(double_commutant : set.centralizer (set.centralizer carrier) = carrier) +(centralizer_centralizer' : + set.centralizer (set.centralizer carrier) = carrier) /-- Consider a von Neumann algebra acting on a Hilbert space `H` as a *-subalgebra of `H →L[ℂ] H`. @@ -77,9 +78,45 @@ or equivalently that it is closed in the weak and strong operator topologies.) add_decl_doc von_neumann_algebra.to_star_subalgebra namespace von_neumann_algebra -variables (H : Type u) [inner_product_space ℂ H] [complete_space H] +variables {H : Type u} [inner_product_space ℂ H] [complete_space H] instance : set_like (von_neumann_algebra H) (H →L[ℂ] H) := -⟨von_neumann_algebra.carrier, λ p q h, by cases p; cases q; congr'⟩ +⟨von_neumann_algebra.carrier, λ S T h, by cases S; cases T; congr'⟩ + +instance : star_mem_class (von_neumann_algebra H) (H →L[ℂ] H) := +{ star_mem := λ s a, s.star_mem' } + +instance : subring_class (von_neumann_algebra H) (H →L[ℂ] H) := +{ add_mem := add_mem', + mul_mem := mul_mem', + one_mem := one_mem', + zero_mem := zero_mem' , + neg_mem := λ s a ha, show -a ∈ s.to_star_subalgebra, from neg_mem ha } + +@[simp] lemma mem_carrier {S : von_neumann_algebra H} {x : H →L[ℂ] H}: + x ∈ S.carrier ↔ x ∈ (S : set (H →L[ℂ] H)) := iff.rfl + +@[ext] theorem ext {S T : von_neumann_algebra H} (h : ∀ x, x ∈ S ↔ x ∈ T) : S = T := +set_like.ext h + +@[simp] lemma centralizer_centralizer (S : von_neumann_algebra H) : + set.centralizer (set.centralizer (S : set (H →L[ℂ] H))) = S := S.centralizer_centralizer' + +/-- The centralizer of a `von_neumann_algebra`, as a `von_neumann_algebra`.-/ +def commutant (S : von_neumann_algebra H) : von_neumann_algebra H := +{ carrier := set.centralizer (S : set (H →L[ℂ] H)), + centralizer_centralizer' := by rw S.centralizer_centralizer, + .. star_subalgebra.centralizer ℂ (S : set (H →L[ℂ] H)) (λ a (ha : a ∈ S), (star_mem ha : _)) } + +@[simp] lemma coe_commutant (S : von_neumann_algebra H) : + ↑S.commutant = set.centralizer (S : set (H →L[ℂ] H)) := rfl + +@[simp] lemma mem_commutant_iff {S : von_neumann_algebra H} {z : H →L[ℂ] H} : + z ∈ S.commutant ↔ ∀ g ∈ S, g * z = z * g := +iff.rfl + +@[simp] lemma commutant_commutant (S : von_neumann_algebra H) : + S.commutant.commutant = S := +set_like.coe_injective S.centralizer_centralizer' end von_neumann_algebra From 1f0096e6caa61e9c849ec2adbd227e960e9dff58 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sun, 22 Jan 2023 18:11:40 +0000 Subject: [PATCH 0326/1291] refactor(data/set/finite): reduce imports (#18245) * Add `eq_or_eq_or_eq_of_forall_not_lt_lt`, `finite.of_forall_not_lt_lt`, `set.finite_of_forall_not_lt_lt` (replacing `set.finite_of_forall_between_eq_endpoints`), and `set.finite_of_forall_not_lt_lt'`. * Import `data.finset.basic` instead of `data.finset.sort` in `data.set.finite`. * Forward-ported in leanprover-community/mathlib4#1738 --- src/data/polynomial/basic.lean | 1 + src/data/set/finite.lean | 46 +++++++++---------- .../constructions/borel_space.lean | 6 +-- src/measure_theory/measure/lebesgue.lean | 14 +----- src/order/basic.lean | 10 ++++ 5 files changed, 37 insertions(+), 40 deletions(-) diff --git a/src/data/polynomial/basic.lean b/src/data/polynomial/basic.lean index 1bd981b8111b8..7b8027fd78cf6 100644 --- a/src/data/polynomial/basic.lean +++ b/src/data/polynomial/basic.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes, Johannes Hölzl, Scott Morrison, Jens Wagemaker -/ import algebra.monoid_algebra.basic +import data.finset.sort /-! # Theory of univariate polynomials diff --git a/src/data/set/finite.lean b/src/data/set/finite.lean index 0cd1144cbf2ce..2c444b6b90190 100644 --- a/src/data/set/finite.lean +++ b/src/data/set/finite.lean @@ -3,7 +3,7 @@ Copyright (c) 2017 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro, Kyle Miller -/ -import data.finset.sort +import data.finset.basic import data.set.functor import data.finite.basic @@ -1227,27 +1227,27 @@ s.finite_to_set.bdd_below end finset -/-- -If a set `s` does not contain any elements between any pair of elements `x, z ∈ s` with `x ≤ z` -(i.e if given `x, y, z ∈ s` such that `x ≤ y ≤ z`, then `y` is either `x` or `z`), then `s` is -finite. --/ -lemma set.finite_of_forall_between_eq_endpoints {α : Type*} [linear_order α] (s : set α) - (h : ∀ (x ∈ s) (y ∈ s) (z ∈ s), x ≤ y → y ≤ z → x = y ∨ y = z) : - set.finite s := +variables [linear_order α] + +/-- If a linear order does not contain any triple of elements `x < y < z`, then this type +is finite. -/ +lemma finite.of_forall_not_lt_lt (h : ∀ ⦃x y z : α⦄, x < y → y < z → false) : + finite α := begin - by_contra hinf, - change s.infinite at hinf, - rcases hinf.exists_subset_card_eq 3 with ⟨t, hts, ht⟩, - let f := t.order_iso_of_fin ht, - let x := f 0, - let y := f 1, - let z := f 2, - have := h x (hts x.2) y (hts y.2) z (hts z.2) - (f.monotone $ by dec_trivial) (f.monotone $ by dec_trivial), - have key₁ : (0 : fin 3) ≠ 1 := by dec_trivial, - have key₂ : (1 : fin 3) ≠ 2 := by dec_trivial, - cases this, - { dsimp only [x, y] at this, exact key₁ (f.injective $ subtype.coe_injective this) }, - { dsimp only [y, z] at this, exact key₂ (f.injective $ subtype.coe_injective this) } + nontriviality α, + rcases exists_pair_ne α with ⟨x, y, hne⟩, + refine @finite.of_fintype α ⟨{x, y}, λ z , _⟩, + simpa [hne] using eq_or_eq_or_eq_of_forall_not_lt_lt h z x y end + +/-- If a set `s` does not contain any triple of elements `x < y < z`, then `s` is finite. -/ +lemma set.finite_of_forall_not_lt_lt {s : set α} (h : ∀ (x y z ∈ s), x < y → y < z → false) : + set.finite s := +@set.to_finite _ s $ finite.of_forall_not_lt_lt $ by simpa only [set_coe.forall'] using h + +lemma set.finite_diff_Union_Ioo (s : set α) : (s \ ⋃ (x ∈ s) (y ∈ s), Ioo x y).finite := +set.finite_of_forall_not_lt_lt $ λ x hx y hy z hz hxy hyz, hy.2 $ mem_Union₂_of_mem hx.1 $ + mem_Union₂_of_mem hz.1 ⟨hxy, hyz⟩ + +lemma set.finite_diff_Union_Ioo' (s : set α) : (s \ ⋃ x : s × s, Ioo x.1 x.2).finite := +by simpa only [Union, supr_prod, supr_subtype] using s.finite_diff_Union_Ioo diff --git a/src/measure_theory/constructions/borel_space.lean b/src/measure_theory/constructions/borel_space.lean index ae3b12da0eabf..4d7319d3a66ad 100644 --- a/src/measure_theory/constructions/borel_space.lean +++ b/src/measure_theory/constructions/borel_space.lean @@ -485,11 +485,7 @@ begin let u := ⋃ (x ∈ s) (y ∈ s), Ioo x y, have huopen : is_open u := is_open_bUnion (λ x hx, is_open_bUnion (λ y hy, is_open_Ioo)), have humeas : measurable_set u := huopen.measurable_set, - have hfinite : (s \ u).finite, - { refine set.finite_of_forall_between_eq_endpoints (s \ u) (λ x hx y hy z hz hxy hyz, _), - by_contra' h, - exact hy.2 (mem_Union₂.mpr ⟨x, hx.1, - mem_Union₂.mpr ⟨z, hz.1, lt_of_le_of_ne hxy h.1, lt_of_le_of_ne hyz h.2⟩⟩) }, + have hfinite : (s \ u).finite := s.finite_diff_Union_Ioo, have : u ⊆ s := Union₂_subset (λ x hx, Union₂_subset (λ y hy, Ioo_subset_Icc_self.trans (h.out hx hy))), rw ← union_diff_cancel this, diff --git a/src/measure_theory/measure/lebesgue.lean b/src/measure_theory/measure/lebesgue.lean index b69c07450b49a..8bd9fe374a079 100644 --- a/src/measure_theory/measure/lebesgue.lean +++ b/src/measure_theory/measure/lebesgue.lean @@ -542,12 +542,7 @@ begin two endpoints, which don't matter since `μ` does not have any atom). -/ let T : s × s → set ℝ := λ p, Ioo p.1 p.2, let u := ⋃ (i : ↥s × ↥s), T i, - have hfinite : (s \ u).finite, - { refine set.finite_of_forall_between_eq_endpoints (s \ u) (λ x hx y hy z hz hxy hyz, _), - by_contra' h, - apply hy.2, - exact mem_Union_of_mem (⟨x, hx.1⟩, ⟨z, hz.1⟩) - ⟨lt_of_le_of_ne hxy h.1, lt_of_le_of_ne hyz h.2⟩ }, + have hfinite : (s \ u).finite := s.finite_diff_Union_Ioo', obtain ⟨A, A_count, hA⟩ : ∃ (A : set (↥s × ↥s)), A.countable ∧ (⋃ (i ∈ A), T i) = ⋃ (i : ↥s × ↥s), T i := is_open_Union_countable _ (λ p, is_open_Ioo), @@ -584,12 +579,7 @@ begin two endpoints, which don't matter since `μ` does not have any atom). -/ let T : s × s → set ℝ := λ p, Ioo p.1 p.2, let u := ⋃ (i : ↥s × ↥s), T i, - have hfinite : (s \ u).finite, - { refine set.finite_of_forall_between_eq_endpoints (s \ u) (λ x hx y hy z hz hxy hyz, _), - by_contra' h, - apply hy.2, - exact mem_Union_of_mem (⟨x, hx.1⟩, ⟨z, hz.1⟩) - ⟨lt_of_le_of_ne hxy h.1, lt_of_le_of_ne hyz h.2⟩ }, + have hfinite : (s \ u).finite := s.finite_diff_Union_Ioo', obtain ⟨A, A_count, hA⟩ : ∃ (A : set (↥s × ↥s)), A.countable ∧ (⋃ (i ∈ A), T i) = ⋃ (i : ↥s × ↥s), T i := is_open_Union_countable _ (λ p, is_open_Ioo), diff --git a/src/order/basic.lean b/src/order/basic.lean index 7d174e19e8aff..61feee4706316 100644 --- a/src/order/basic.lean +++ b/src/order/basic.lean @@ -876,6 +876,16 @@ or_iff_not_imp_left.2 $ λ h, ⟨λ a ha₁, le_of_not_gt $ λ ha₂, h ⟨a, ha₁, ha₂⟩, λ a ha₂, le_of_not_gt $ λ ha₁, h ⟨a, ha₁, ha₂⟩⟩ +/-- If a linear order has no elements `x < y < z`, then it has at most two elements. -/ +lemma eq_or_eq_or_eq_of_forall_not_lt_lt {α : Type*} [linear_order α] + (h : ∀ ⦃x y z : α⦄, x < y → y < z → false) (x y z : α) : x = y ∨ y = z ∨ x = z := +begin + by_contra hne, push_neg at hne, + cases hne.1.lt_or_lt with h₁ h₁; cases hne.2.1.lt_or_lt with h₂ h₂; + cases hne.2.2.lt_or_lt with h₃ h₃, + exacts [h h₁ h₂, h h₂ h₃, h h₃ h₂, h h₃ h₁, h h₁ h₃, h h₂ h₃, h h₁ h₃, h h₂ h₁] +end + namespace punit variables (a b : punit.{u+1}) From 2032a878972d5672e7c27c957e7a6e297b044973 Mon Sep 17 00:00:00 2001 From: Paul Lezeau <72892199+Paul-Lez@users.noreply.github.com> Date: Mon, 23 Jan 2023 08:58:58 +0000 Subject: [PATCH 0327/1291] chore(field_theory/minpoly/*): replace `gcd_monoid.lean` by `is_integrally_closed.lean` and remove results that have been generalized (#18206) --- src/algebra/gcd_monoid/integrally_closed.lean | 2 +- src/field_theory/minpoly/default.lean | 2 +- ..._monoid.lean => is_integrally_closed.lean} | 145 ++--------- .../cyclotomic/discriminant.lean | 4 +- src/number_theory/cyclotomic/rat.lean | 4 +- .../number_field/embeddings.lean | 4 +- .../dedekind_domain/adic_valuation.lean | 1 + .../dedekind_domain/selmer_group.lean | 1 + src/ring_theory/is_adjoin_root.lean | 6 +- .../polynomial/cyclotomic/basic.lean | 7 +- .../polynomial/eisenstein/basic.lean | 233 ++++++++++++++++++ .../is_integral.lean} | 226 +---------------- src/ring_theory/roots_of_unity.lean | 20 +- 13 files changed, 289 insertions(+), 366 deletions(-) rename src/field_theory/minpoly/{gcd_monoid.lean => is_integrally_closed.lean} (53%) create mode 100644 src/ring_theory/polynomial/eisenstein/basic.lean rename src/ring_theory/polynomial/{eisenstein.lean => eisenstein/is_integral.lean} (68%) diff --git a/src/algebra/gcd_monoid/integrally_closed.lean b/src/algebra/gcd_monoid/integrally_closed.lean index 341d93937c36b..47ff732f99523 100644 --- a/src/algebra/gcd_monoid/integrally_closed.lean +++ b/src/algebra/gcd_monoid/integrally_closed.lean @@ -5,7 +5,7 @@ Authors: Andrew Yang -/ import algebra.gcd_monoid.basic import ring_theory.integrally_closed -import ring_theory.polynomial.eisenstein +import ring_theory.polynomial.eisenstein.basic /-! diff --git a/src/field_theory/minpoly/default.lean b/src/field_theory/minpoly/default.lean index aa4d2da646584..e7bc53a8783e6 100644 --- a/src/field_theory/minpoly/default.lean +++ b/src/field_theory/minpoly/default.lean @@ -1,3 +1,3 @@ import field_theory.minpoly.basic import field_theory.minpoly.field -import field_theory.minpoly.gcd_monoid +import field_theory.minpoly.is_integrally_closed diff --git a/src/field_theory/minpoly/gcd_monoid.lean b/src/field_theory/minpoly/is_integrally_closed.lean similarity index 53% rename from src/field_theory/minpoly/gcd_monoid.lean rename to src/field_theory/minpoly/is_integrally_closed.lean index 5f14cb650f03c..37c45db1fdf8e 100644 --- a/src/field_theory/minpoly/gcd_monoid.lean +++ b/src/field_theory/minpoly/is_integrally_closed.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2019 Riccardo Brasca. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Riccardo Brasca +Authors: Riccardo Brasca, Paul Lezeau, Junyan Xu -/ import data.polynomial.field_division import ring_theory.adjoin_root @@ -15,23 +15,16 @@ This file specializes the theory of minpoly to the case of an algebra over a GCD ## Main results - * `gcd_domain_eq_field_fractions`: For GCD domains, the minimal polynomial over the ring is the - same as the minimal polynomial over the fraction field. - - * `gcd_domain_dvd` : For GCD domains, the minimal polynomial divides any primitive polynomial - that has the integral element as root. - - * `gcd_domain_unique` : The minimal polynomial of an element `x` is uniquely characterized by - its defining property: if there is another monic polynomial of minimal degree that has `x` as a - root, then this polynomial is equal to the minimal polynomial of `x`. - * `is_integrally_closed_eq_field_fractions`: For integrally closed domains, the minimal polynomial over the ring is the same as the minimal polynomial over the fraction field. -## Todo + * `is_integrally_closed_dvd` : For integrally closed domains, the minimal polynomial divides any + primitive polynomial that has the integral element as root. + + * `is_integrally_closed_unique` : The minimal polynomial of an element `x` is uniquely + characterized by its defining property: if there is another monic polynomial of minimal degree + that has `x` as a root, then this polynomial is equal to the minimal polynomial of `x`. - * Remove all results that are now special cases (e.g. we no longer need `gcd_monoid_dvd` since we - have `is_integrally_closed_dvd`). -/ open_locale classical polynomial @@ -41,43 +34,11 @@ namespace minpoly variables {R S : Type*} [comm_ring R] [comm_ring S] [is_domain R] [algebra R S] - section variables (K L : Type*) [field K] [algebra R K] [is_fraction_ring R K] [field L] [algebra R L] [algebra S L] [algebra K L] [is_scalar_tower R K L] [is_scalar_tower R S L] -section gcd_domain - -variable [normalized_gcd_monoid R] - -/-- For GCD domains, the minimal polynomial over the ring is the same as the minimal polynomial -over the fraction field. See `minpoly.gcd_domain_eq_field_fractions'` if `S` is already a -`K`-algebra. -/ -lemma gcd_domain_eq_field_fractions [is_domain S] {s : S} (hs : is_integral R s) : - minpoly K (algebra_map S L s) = (minpoly R s).map (algebra_map R K) := -begin - refine (eq_of_irreducible_of_monic _ _ _).symm, - { exact (polynomial.is_primitive.irreducible_iff_irreducible_map_fraction_map - (polynomial.monic.is_primitive (monic hs))).1 (irreducible hs) }, - { rw [aeval_map_algebra_map, aeval_algebra_map_apply, aeval, map_zero] }, - { exact (monic hs).map _ } -end - -/-- For GCD domains, the minimal polynomial over the ring is the same as the minimal polynomial -over the fraction field. Compared to `minpoly.gcd_domain_eq_field_fractions`, this version is useful -if the element is in a ring that is already a `K`-algebra. -/ -lemma gcd_domain_eq_field_fractions' [is_domain S] [algebra K S] [is_scalar_tower R K S] - {s : S} (hs : is_integral R s) : minpoly K s = (minpoly R s).map (algebra_map R K) := -begin - let L := fraction_ring S, - rw [← gcd_domain_eq_field_fractions K L hs], - refine minpoly.eq_of_algebra_map_eq (is_fraction_ring.injective S L) - (is_integral_of_is_scalar_tower hs) rfl -end - -end gcd_domain - variables [is_integrally_closed R] /-- For integrally closed domains, the minimal polynomial over the ring is the same as the minimal @@ -105,76 +66,19 @@ begin (is_integral_of_is_scalar_tower hs) rfl end -/-- For GCD domains, the minimal polynomial over the ring is the same as the minimal polynomial -over the fraction field. Compared to `minpoly.is_integrally_closed_eq_field_fractions`, this -version is useful if the element is in a ring that is not a domain -/ -theorem is_integrally_closed_eq_field_fractions'' [no_zero_smul_divisors S L] {s : S} - (hs : is_integral R s) : minpoly K (algebra_map S L s) = map (algebra_map R K) (minpoly R s) := -begin - --the idea of the proof is the following: since the minpoly of `a` over `Frac(R)` divides the - --minpoly of `a` over `R`, it is itself in `R`. Hence its degree is greater or equal to that of - --the minpoly of `a` over `R`. But the minpoly of `a` over `Frac(R)` divides the minpoly of a - --over `R` in `R[X]` so we are done. - - --a few "trivial" preliminary results to set up the proof - have lem0 : minpoly K (algebra_map S L s) ∣ (map (algebra_map R K) (minpoly R s)), - { exact dvd_map_of_is_scalar_tower' R K L s }, - - have lem1 : is_integral K (algebra_map S L s), - { refine is_integral_map_of_comp_eq_of_is_integral (algebra_map R K) _ _ hs, - rw [← is_scalar_tower.algebra_map_eq, ← is_scalar_tower.algebra_map_eq] }, - - obtain ⟨g, hg⟩ := is_integrally_closed.eq_map_mul_C_of_dvd K (minpoly.monic hs) lem0, - rw [(minpoly.monic lem1).leading_coeff, C_1, mul_one] at hg, - have lem2 : polynomial.aeval s g = 0, - { have := minpoly.aeval K (algebra_map S L s), - rw [← hg, ← map_aeval_eq_aeval_map, ← map_zero (algebra_map S L)] at this, - { exact no_zero_smul_divisors.algebra_map_injective S L this }, - { rw [← is_scalar_tower.algebra_map_eq, ← is_scalar_tower.algebra_map_eq] } }, - - have lem3 : g.monic, - { simpa only [function.injective.monic_map_iff (is_fraction_ring.injective R K), hg] - using minpoly.monic lem1 }, - - rw [← hg], - refine congr_arg _ (eq.symm (polynomial.eq_of_monic_of_dvd_of_nat_degree_le lem3 - (minpoly.monic hs) _ _)), - { rwa [← map_dvd_map _ (is_fraction_ring.injective R K) lem3, hg] }, - { exact nat_degree_le_nat_degree (minpoly.min R s lem3 lem2) }, -end - end variables [is_domain S] [no_zero_smul_divisors R S] -lemma gcd_domain_dvd [normalized_gcd_monoid R] {s : S} (hs : is_integral R s) {P : R[X]} - (hP : P ≠ 0) (hroot : polynomial.aeval s P = 0) : minpoly R s ∣ P := -begin - let K := fraction_ring R, - let L := fraction_ring S, - let P₁ := P.prim_part, - suffices : minpoly R s ∣ P₁, - { exact dvd_trans this (prim_part_dvd _) }, - apply (is_primitive.dvd_iff_fraction_map_dvd_fraction_map K (monic hs).is_primitive - P.is_primitive_prim_part).2, - let y := algebra_map S L s, - have hy : is_integral R y := hs.algebra_map, - rw [← gcd_domain_eq_field_fractions K L hs], - refine dvd _ _ _, - rw [aeval_map_algebra_map, aeval_algebra_map_apply, aeval_prim_part_eq_zero hP hroot, map_zero] -end - variable [is_integrally_closed R] /-- For integrally closed rings, the minimal polynomial divides any polynomial that has the integral element as root. See also `minpoly.dvd` which relaxes the assumptions on `S` in exchange for stronger assumptions on `R`. -/ -theorem is_integrally_closed_dvd [nontrivial R] (p : R[X]) {s : S} (hs : is_integral R s) : - polynomial.aeval s p = 0 ↔ minpoly R s ∣ p := +theorem is_integrally_closed_dvd [nontrivial R] {s : S} (hs : is_integral R s) {p : R[X]} + (hp : polynomial.aeval s p = 0) : minpoly R s ∣ p := begin - refine ⟨λ hp, _, λ hp, _⟩, - - { let K := fraction_ring R, + let K := fraction_ring R, let L := fraction_ring S, have : minpoly K (algebra_map S L s) ∣ map (algebra_map R K) (p %ₘ (minpoly R s)), { rw [map_mod_by_monic _ (minpoly.monic hs), mod_by_monic_eq_sub_mul_div], @@ -183,30 +87,27 @@ begin rw [← is_scalar_tower.algebra_map_eq, ← is_scalar_tower.algebra_map_eq], apply dvd_mul_of_dvd_left, - rw is_integrally_closed_eq_field_fractions'' K L hs, + rw is_integrally_closed_eq_field_fractions K L hs, exact monic.map _ (minpoly.monic hs) }, - rw [is_integrally_closed_eq_field_fractions'' _ _ hs, map_dvd_map (algebra_map R K) + rw [is_integrally_closed_eq_field_fractions _ _ hs, map_dvd_map (algebra_map R K) (is_fraction_ring.injective R K) (minpoly.monic hs)] at this, rw [← dvd_iff_mod_by_monic_eq_zero (minpoly.monic hs)], refine polynomial.eq_zero_of_dvd_of_degree_lt this (degree_mod_by_monic_lt p $ minpoly.monic hs), - all_goals { apply_instance } }, - - { simpa only [ring_hom.mem_ker, ring_hom.coe_comp, coe_eval_ring_hom, - coe_map_ring_hom, function.comp_app, eval_map, ← aeval_def] using - aeval_eq_zero_of_dvd_aeval_eq_zero hp (minpoly.aeval R s) } + all_goals { apply_instance } end +theorem is_integrally_closed_dvd_iff [nontrivial R] {s : S} (hs : is_integral R s) (p : R[X]) : + polynomial.aeval s p = 0 ↔ minpoly R s ∣ p := +⟨λ hp, is_integrally_closed_dvd hs hp, λ hp, by simpa only [ring_hom.mem_ker, ring_hom.coe_comp, + coe_eval_ring_hom, coe_map_ring_hom, function.comp_app, eval_map, ← aeval_def] using + aeval_eq_zero_of_dvd_aeval_eq_zero hp (minpoly.aeval R s)⟩ + lemma ker_eval {s : S} (hs : is_integral R s) : ((polynomial.aeval s).to_ring_hom : R[X] →+* S).ker = ideal.span ({minpoly R s} : set R[X] ):= -begin - apply le_antisymm; intros p hp, - { rwa [ring_hom.mem_ker, alg_hom.to_ring_hom_eq_coe, alg_hom.coe_to_ring_hom, - is_integrally_closed_dvd p hs, ← ideal.mem_span_singleton] at hp }, - { rwa [ring_hom.mem_ker, alg_hom.to_ring_hom_eq_coe, alg_hom.coe_to_ring_hom, - is_integrally_closed_dvd p hs, ← ideal.mem_span_singleton] } -end +by ext p ; simp_rw [ring_hom.mem_ker, alg_hom.to_ring_hom_eq_coe, alg_hom.coe_to_ring_hom, + is_integrally_closed_dvd_iff hs, ← ideal.mem_span_singleton] /-- If an element `x` is a root of a nonzero polynomial `p`, then the degree of `p` is at least the degree of the minimal polynomial of `x`. See also `minpoly.degree_le_of_ne_zero` which relaxes the @@ -216,7 +117,7 @@ lemma is_integrally_closed.degree_le_of_ne_zero {s : S} (hs : is_integral R s) { begin rw [degree_eq_nat_degree (minpoly.ne_zero hs), degree_eq_nat_degree hp0], norm_cast, - exact nat_degree_le_of_dvd ((is_integrally_closed_dvd _ hs).mp hp) hp0 + exact nat_degree_le_of_dvd ((is_integrally_closed_dvd_iff hs _).mp hp) hp0 end /-- The minimal polynomial of an element `x` is uniquely characterized by its defining property: @@ -255,7 +156,7 @@ begin by_cases hPzero : P = 0, { simpa [hPzero] using hP.symm }, rw [← hP, minpoly.to_adjoin_apply', lift_hom_mk, ← subalgebra.coe_eq_zero, - aeval_subalgebra_coe, set_like.coe_mk, is_integrally_closed_dvd _ hx] at hP₁, + aeval_subalgebra_coe, set_like.coe_mk, is_integrally_closed_dvd_iff hx] at hP₁, obtain ⟨Q, hQ⟩ := hP₁, rw [← hP, hQ, ring_hom.map_mul, mk_self, zero_mul], end diff --git a/src/number_theory/cyclotomic/discriminant.lean b/src/number_theory/cyclotomic/discriminant.lean index f268b04516fcd..2a254dc3a5519 100644 --- a/src/number_theory/cyclotomic/discriminant.lean +++ b/src/number_theory/cyclotomic/discriminant.lean @@ -44,9 +44,9 @@ begin (λ i j, to_matrix_is_integral H₁ _ _ _ _) (λ i j, to_matrix_is_integral H₂ _ _ _ _), { exact hζ.is_integral n.pos }, - { refine minpoly.gcd_domain_eq_field_fractions' _ (hζ.is_integral n.pos) }, + { refine minpoly.is_integrally_closed_eq_field_fractions' _ (hζ.is_integral n.pos) }, { exact is_integral_sub (hζ.is_integral n.pos) is_integral_one }, - { refine minpoly.gcd_domain_eq_field_fractions' _ _, + { refine minpoly.is_integrally_closed_eq_field_fractions' _ _, exact is_integral_sub (hζ.is_integral n.pos) is_integral_one } end diff --git a/src/number_theory/cyclotomic/rat.lean b/src/number_theory/cyclotomic/rat.lean index 67def7f2ff0c1..b7f1fdd822374 100644 --- a/src/number_theory/cyclotomic/rat.lean +++ b/src/number_theory/cyclotomic/rat.lean @@ -5,7 +5,7 @@ Authors: Riccardo Brasca -/ import number_theory.cyclotomic.discriminant -import ring_theory.polynomial.eisenstein +import ring_theory.polynomial.eisenstein.is_integral /-! # Ring of integers of `p ^ n`-th cyclotomic fields @@ -105,7 +105,7 @@ begin 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₁ := minpoly.is_integrally_closed_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₁, diff --git a/src/number_theory/number_field/embeddings.lean b/src/number_theory/number_field/embeddings.lean index 3e5e993c7678b..def8c20de89d5 100644 --- a/src/number_theory/number_field/embeddings.lean +++ b/src/number_theory/number_field/embeddings.lean @@ -6,7 +6,7 @@ Authors: Alex J. Best, Xavier Roblot import analysis.complex.polynomial import data.complex.basic -import field_theory.minpoly.gcd_monoid +import field_theory.minpoly.is_integrally_closed import number_theory.number_field.basic /-! @@ -101,7 +101,7 @@ begin let C := nat.ceil ((max B 1) ^ (finrank ℚ K) * (finrank ℚ K).choose ((finrank ℚ K) / 2)), have := bUnion_roots_finite (algebra_map ℤ K) (finrank ℚ K) (finite_Icc (-C : ℤ) C), refine this.subset (λ x hx, _), simp_rw mem_Union, - have h_map_ℚ_minpoly := minpoly.gcd_domain_eq_field_fractions' ℚ hx.1, + have h_map_ℚ_minpoly := minpoly.is_integrally_closed_eq_field_fractions' ℚ hx.1, refine ⟨_, ⟨_, λ i, _⟩, mem_root_set.2 ⟨minpoly.ne_zero hx.1, minpoly.aeval ℤ x⟩⟩, { rw [← (minpoly.monic hx.1).nat_degree_map (algebra_map ℤ ℚ), ← h_map_ℚ_minpoly], exact minpoly.nat_degree_le (is_integral_of_is_scalar_tower hx.1) }, diff --git a/src/ring_theory/dedekind_domain/adic_valuation.lean b/src/ring_theory/dedekind_domain/adic_valuation.lean index 5098e9e5faf79..ea551209d81bb 100644 --- a/src/ring_theory/dedekind_domain/adic_valuation.lean +++ b/src/ring_theory/dedekind_domain/adic_valuation.lean @@ -6,6 +6,7 @@ Authors: María Inés de Frutos-Fernández import ring_theory.dedekind_domain.ideal import ring_theory.valuation.extend_to_localization import ring_theory.valuation.valuation_subring +import ring_theory.polynomial.cyclotomic.basic import topology.algebra.valued_field /-! diff --git a/src/ring_theory/dedekind_domain/selmer_group.lean b/src/ring_theory/dedekind_domain/selmer_group.lean index 433814c96a50e..f729e75d1c898 100644 --- a/src/ring_theory/dedekind_domain/selmer_group.lean +++ b/src/ring_theory/dedekind_domain/selmer_group.lean @@ -6,6 +6,7 @@ Authors: David Kurniadi Angdinata import algebra.hom.equiv.type_tags import data.zmod.quotient import ring_theory.dedekind_domain.adic_valuation +import ring_theory.norm /-! # Selmer groups of fraction fields of Dedekind domains diff --git a/src/ring_theory/is_adjoin_root.lean b/src/ring_theory/is_adjoin_root.lean index a9d74b7173734..766e7a34bfebb 100644 --- a/src/ring_theory/is_adjoin_root.lean +++ b/src/ring_theory/is_adjoin_root.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Anne Baanen -/ import data.polynomial.algebra_map -import field_theory.minpoly.gcd_monoid +import field_theory.minpoly.is_integrally_closed import ring_theory.power_basis /-! @@ -627,10 +627,10 @@ end is_adjoin_root namespace is_adjoin_root_monic -lemma minpoly_eq [is_domain R] [is_domain S] [no_zero_smul_divisors R S] [normalized_gcd_monoid R] +lemma minpoly_eq [is_domain R] [is_domain S] [no_zero_smul_divisors R S] [is_integrally_closed R] (h : is_adjoin_root_monic S f) (hirr : irreducible f) : minpoly R h.root = f := -let ⟨q, hq⟩ := minpoly.gcd_domain_dvd h.is_integral_root h.monic.ne_zero h.aeval_root in +let ⟨q, hq⟩ := minpoly.is_integrally_closed_dvd h.is_integral_root h.aeval_root in symm $ eq_of_monic_of_associated h.monic (minpoly.monic h.is_integral_root) $ by convert (associated.mul_left (minpoly R h.root) $ associated_one_iff_is_unit.2 $ (hirr.is_unit_or_is_unit hq).resolve_left $ diff --git a/src/ring_theory/polynomial/cyclotomic/basic.lean b/src/ring_theory/polynomial/cyclotomic/basic.lean index 141cc8845711e..90332b2db71e3 100644 --- a/src/ring_theory/polynomial/cyclotomic/basic.lean +++ b/src/ring_theory/polynomial/cyclotomic/basic.lean @@ -787,7 +787,7 @@ lemma _root_.is_primitive_root.minpoly_dvd_cyclotomic {n : ℕ} {K : Type*} [fie (h : is_primitive_root μ n) (hpos : 0 < n) [char_zero K] : minpoly ℤ μ ∣ cyclotomic n ℤ := begin - apply minpoly.gcd_domain_dvd (is_integral h hpos) (cyclotomic_ne_zero n ℤ), + apply minpoly.is_integrally_closed_dvd (is_integral h hpos), simpa [aeval_def, eval₂_eq_eval_map, is_root.def] using is_root_cyclotomic hpos h end @@ -816,7 +816,7 @@ lemma cyclotomic_eq_minpoly_rat {n : ℕ} {K : Type*} [field K] {μ : K} cyclotomic n ℚ = minpoly ℚ μ := begin rw [← map_cyclotomic_int, cyclotomic_eq_minpoly h hpos], - exact (minpoly.gcd_domain_eq_field_fractions' _ (is_integral h hpos)).symm + exact (minpoly.is_integrally_closed_eq_field_fractions' _ (is_integral h hpos)).symm end /-- `cyclotomic n ℤ` is irreducible. -/ @@ -909,8 +909,7 @@ begin { have hpos := nat.mul_pos hzero hp.pos, have hprim := complex.is_primitive_root_exp _ hpos.ne.symm, rw [cyclotomic_eq_minpoly hprim hpos], - refine minpoly.gcd_domain_dvd (hprim.is_integral hpos) - ((cyclotomic.monic n ℤ).expand hp.pos).ne_zero _, + refine minpoly.is_integrally_closed_dvd (hprim.is_integral hpos) _, rw [aeval_def, ← eval_map, map_expand, map_cyclotomic, expand_eval, ← is_root.def, is_root_cyclotomic_iff], { convert is_primitive_root.pow_of_dvd hprim hp.ne_zero (dvd_mul_left p n), diff --git a/src/ring_theory/polynomial/eisenstein/basic.lean b/src/ring_theory/polynomial/eisenstein/basic.lean new file mode 100644 index 0000000000000..6d37190e8da92 --- /dev/null +++ b/src/ring_theory/polynomial/eisenstein/basic.lean @@ -0,0 +1,233 @@ +/- +Copyright (c) 2022 Riccardo Brasca. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Riccardo Brasca +-/ + +import ring_theory.eisenstein_criterion +import ring_theory.polynomial.scale_roots + +/-! +# Eisenstein polynomials +Given an ideal `𝓟` of a commutative semiring `R`, we say that a polynomial `f : R[X]` is +*Eisenstein at `𝓟`* if `f.leading_coeff ∉ 𝓟`, `∀ n, n < f.nat_degree → f.coeff n ∈ 𝓟` and +`f.coeff 0 ∉ 𝓟 ^ 2`. In this file we gather miscellaneous results about Eisenstein polynomials. + +## Main definitions +* `polynomial.is_eisenstein_at f 𝓟`: the property of being Eisenstein at `𝓟`. + +## Main results +* `polynomial.is_eisenstein_at.irreducible`: if a primitive `f` satisfies `f.is_eisenstein_at 𝓟`, + where `𝓟.is_prime`, then `f` is irreducible. + +## Implementation details +We also define a notion `is_weakly_eisenstein_at` requiring only that +`∀ n < f.nat_degree → f.coeff n ∈ 𝓟`. This makes certain results slightly more general and it is +useful since it is sometimes better behaved (for example it is stable under `polynomial.map`). + +-/ + +universes u v w z + +variables {R : Type u} + +open ideal algebra finset + +open_locale big_operators polynomial + +namespace polynomial + +/-- Given an ideal `𝓟` of a commutative semiring `R`, we say that a polynomial `f : R[X]` +is *weakly Eisenstein at `𝓟`* if `∀ n, n < f.nat_degree → f.coeff n ∈ 𝓟`. -/ +@[mk_iff] structure is_weakly_eisenstein_at [comm_semiring R] (f : R[X]) (𝓟 : ideal R) : + Prop := (mem : ∀ {n}, n < f.nat_degree → f.coeff n ∈ 𝓟) + +/-- Given an ideal `𝓟` of a commutative semiring `R`, we say that a polynomial `f : R[X]` +is *Eisenstein at `𝓟`* if `f.leading_coeff ∉ 𝓟`, `∀ n, n < f.nat_degree → f.coeff n ∈ 𝓟` and +`f.coeff 0 ∉ 𝓟 ^ 2`. -/ +@[mk_iff] structure is_eisenstein_at [comm_semiring R] (f : R[X]) (𝓟 : ideal R) : Prop := +(leading : f.leading_coeff ∉ 𝓟) +(mem : ∀ {n}, n < f.nat_degree → f.coeff n ∈ 𝓟) +(not_mem : f.coeff 0 ∉ 𝓟 ^ 2) + +namespace is_weakly_eisenstein_at + +section comm_semiring + +variables [comm_semiring R] {𝓟 : ideal R} {f : R[X]} (hf : f.is_weakly_eisenstein_at 𝓟) + +include hf + +lemma map {A : Type v} [comm_ring A] (φ : R →+* A) : (f.map φ).is_weakly_eisenstein_at (𝓟.map φ) := +begin + refine (is_weakly_eisenstein_at_iff _ _).2 (λ n hn, _), + rw [coeff_map], + exact mem_map_of_mem _ (hf.mem (lt_of_lt_of_le hn (nat_degree_map_le _ _))) +end + +end comm_semiring + +section comm_ring + +variables [comm_ring R] {𝓟 : ideal R} {f : R[X]} (hf : f.is_weakly_eisenstein_at 𝓟) +variables {S : Type v} [comm_ring S] [algebra R S] + +section principal + +variable {p : R} + +local notation `P` := submodule.span R {p} + +lemma exists_mem_adjoin_mul_eq_pow_nat_degree {x : S} (hx : aeval x f = 0) + (hmo : f.monic) (hf : f.is_weakly_eisenstein_at P) : ∃ y ∈ adjoin R ({x} : set S), + (algebra_map R S) p * y = x ^ (f.map (algebra_map R S)).nat_degree := +begin + rw [aeval_def, polynomial.eval₂_eq_eval_map, eval_eq_sum_range, range_add_one, + sum_insert not_mem_range_self, sum_range, (hmo.map + (algebra_map R S)).coeff_nat_degree, one_mul] at hx, + replace hx := eq_neg_of_add_eq_zero_left hx, + have : ∀ n < f.nat_degree, p ∣ f.coeff n, + { intros n hn, + refine mem_span_singleton.1 (by simpa using hf.mem hn) }, + choose! φ hφ using this, + conv_rhs at hx { congr, congr, skip, funext, + rw [fin.coe_eq_val, coeff_map, hφ i.1 (lt_of_lt_of_le i.2 (nat_degree_map_le _ _)), + ring_hom.map_mul, mul_assoc] }, + rw [hx, ← mul_sum, neg_eq_neg_one_mul, ← mul_assoc (-1 : S), mul_comm (-1 : S), mul_assoc], + refine ⟨-1 * ∑ (i : fin (f.map (algebra_map R S)).nat_degree), + (algebra_map R S) (φ i.1) * x ^ i.1, _, rfl⟩, + exact subalgebra.mul_mem _ (subalgebra.neg_mem _ (subalgebra.one_mem _)) + (subalgebra.sum_mem _ (λ i hi, subalgebra.mul_mem _ (subalgebra.algebra_map_mem _ _) + (subalgebra.pow_mem _ (subset_adjoin (set.mem_singleton x)) _))) +end + +lemma exists_mem_adjoin_mul_eq_pow_nat_degree_le {x : S} (hx : aeval x f = 0) + (hmo : f.monic) (hf : f.is_weakly_eisenstein_at P) : + ∀ i, (f.map (algebra_map R S)).nat_degree ≤ i → + ∃ y ∈ adjoin R ({x} : set S), (algebra_map R S) p * y = x ^ i := +begin + intros i hi, + obtain ⟨k, hk⟩ := exists_add_of_le hi, + rw [hk, pow_add], + obtain ⟨y, hy, H⟩ := exists_mem_adjoin_mul_eq_pow_nat_degree hx hmo hf, + refine ⟨y * x ^ k, _, _⟩, + { exact subalgebra.mul_mem _ hy (subalgebra.pow_mem _ (subset_adjoin (set.mem_singleton x)) _) }, + { rw [← mul_assoc _ y, H] } +end + +end principal + +include hf + +lemma pow_nat_degree_le_of_root_of_monic_mem {x : R} (hroot : is_root f x) (hmo : f.monic) : + ∀ i, f.nat_degree ≤ i → x ^ i ∈ 𝓟 := +begin + intros i hi, + obtain ⟨k, hk⟩ := exists_add_of_le hi, + rw [hk, pow_add], + suffices : x ^ f.nat_degree ∈ 𝓟, + { exact mul_mem_right (x ^ k) 𝓟 this }, + rw [is_root.def, eval_eq_sum_range, finset.range_add_one, finset.sum_insert + finset.not_mem_range_self, finset.sum_range, hmo.coeff_nat_degree, one_mul] at hroot, + rw [eq_neg_of_add_eq_zero_left hroot, neg_mem_iff], + refine submodule.sum_mem _ (λ i hi, mul_mem_right _ _ (hf.mem (fin.is_lt i))) +end + +lemma pow_nat_degree_le_of_aeval_zero_of_monic_mem_map {x : S} (hx : aeval x f = 0) + (hmo : f.monic) : + ∀ i, (f.map (algebra_map R S)).nat_degree ≤ i → x ^ i ∈ 𝓟.map (algebra_map R S) := +begin + suffices : x ^ (f.map (algebra_map R S)).nat_degree ∈ 𝓟.map (algebra_map R S), + { intros i hi, + obtain ⟨k, hk⟩ := exists_add_of_le hi, + rw [hk, pow_add], + refine mul_mem_right _ _ this }, + rw [aeval_def, eval₂_eq_eval_map, ← is_root.def] at hx, + refine pow_nat_degree_le_of_root_of_monic_mem (hf.map _) hx (hmo.map _) _ rfl.le +end + +end comm_ring + +end is_weakly_eisenstein_at + +section scale_roots + +variables {A : Type*} [comm_ring R] [comm_ring A] + +lemma scale_roots.is_weakly_eisenstein_at (p : R[X]) {x : R} + {P : ideal R} (hP : x ∈ P) : (scale_roots p x).is_weakly_eisenstein_at P := +begin + refine ⟨λ i hi, _⟩, + rw coeff_scale_roots, + rw [nat_degree_scale_roots, ← tsub_pos_iff_lt] at hi, + exact ideal.mul_mem_left _ _ (ideal.pow_mem_of_mem P hP _ hi) +end + +lemma dvd_pow_nat_degree_of_eval₂_eq_zero {f : R →+* A} + (hf : function.injective f) {p : R[X]} (hp : p.monic) (x y : R) (z : A) + (h : p.eval₂ f z = 0) (hz : f x * z = f y) : x ∣ y ^ p.nat_degree := +begin + rw [← nat_degree_scale_roots p x, ← ideal.mem_span_singleton], + refine (scale_roots.is_weakly_eisenstein_at _ (ideal.mem_span_singleton.mpr $ dvd_refl x)) + .pow_nat_degree_le_of_root_of_monic_mem _ ((monic_scale_roots_iff x).mpr hp) _ le_rfl, + rw injective_iff_map_eq_zero' at hf, + have := scale_roots_eval₂_eq_zero f h, + rwa [hz, polynomial.eval₂_at_apply, hf] at this +end + +lemma dvd_pow_nat_degree_of_aeval_eq_zero [algebra R A] [nontrivial A] + [no_zero_smul_divisors R A] {p : R[X]} (hp : p.monic) (x y : R) (z : A) + (h : polynomial.aeval z p = 0) (hz : z * algebra_map R A x = algebra_map R A y) : + x ∣ y ^ p.nat_degree := +dvd_pow_nat_degree_of_eval₂_eq_zero (no_zero_smul_divisors.algebra_map_injective R A) + hp x y z h ((mul_comm _ _).trans hz) + +end scale_roots + +namespace is_eisenstein_at + +section comm_semiring + +variables [comm_semiring R] {𝓟 : ideal R} {f : R[X]} (hf : f.is_eisenstein_at 𝓟) + +lemma _root_.polynomial.monic.leading_coeff_not_mem (hf : f.monic) (h : 𝓟 ≠ ⊤) : + ¬f.leading_coeff ∈ 𝓟 := +hf.leading_coeff.symm ▸ (ideal.ne_top_iff_one _).1 h + +lemma _root_.polynomial.monic.is_eisenstein_at_of_mem_of_not_mem (hf : f.monic) (h : 𝓟 ≠ ⊤) + (hmem : ∀ {n}, n < f.nat_degree → f.coeff n ∈ 𝓟) (hnot_mem : f.coeff 0 ∉ 𝓟 ^ 2) : + f.is_eisenstein_at 𝓟 := +{ leading := hf.leading_coeff_not_mem h, + mem := λ n hn, hmem hn, + not_mem := hnot_mem } + +include hf + +lemma is_weakly_eisenstein_at : is_weakly_eisenstein_at f 𝓟 := ⟨λ _, hf.mem⟩ + +lemma coeff_mem {n : ℕ} (hn : n ≠ f.nat_degree) : f.coeff n ∈ 𝓟 := +begin + cases ne_iff_lt_or_gt.1 hn, + { exact hf.mem h }, + { rw [coeff_eq_zero_of_nat_degree_lt h], + exact ideal.zero_mem _} +end + +end comm_semiring + +section is_domain + +variables [comm_ring R] [is_domain R] {𝓟 : ideal R} {f : R[X]} (hf : f.is_eisenstein_at 𝓟) + +/-- If a primitive `f` satisfies `f.is_eisenstein_at 𝓟`, where `𝓟.is_prime`, then `f` is +irreducible. -/ +lemma irreducible (hprime : 𝓟.is_prime) (hu : f.is_primitive) + (hfd0 : 0 < f.nat_degree) : irreducible f := +irreducible_of_eisenstein_criterion hprime hf.leading (λ n hn, hf.mem (coe_lt_degree.1 hn)) + (nat_degree_pos_iff_degree_pos.1 hfd0) hf.not_mem hu + +end is_domain + +end is_eisenstein_at + +end polynomial diff --git a/src/ring_theory/polynomial/eisenstein.lean b/src/ring_theory/polynomial/eisenstein/is_integral.lean similarity index 68% rename from src/ring_theory/polynomial/eisenstein.lean rename to src/ring_theory/polynomial/eisenstein/is_integral.lean index e1c05e1080d69..9287f49e81052 100644 --- a/src/ring_theory/polynomial/eisenstein.lean +++ b/src/ring_theory/polynomial/eisenstein/is_integral.lean @@ -4,23 +4,15 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Riccardo Brasca -/ -import ring_theory.eisenstein_criterion import ring_theory.integrally_closed import ring_theory.norm import ring_theory.polynomial.cyclotomic.basic /-! # Eisenstein polynomials -Given an ideal `𝓟` of a commutative semiring `R`, we say that a polynomial `f : R[X]` is -*Eisenstein at `𝓟`* if `f.leading_coeff ∉ 𝓟`, `∀ n, n < f.nat_degree → f.coeff n ∈ 𝓟` and -`f.coeff 0 ∉ 𝓟 ^ 2`. In this file we gather miscellaneous results about Eisenstein polynomials. - -## Main definitions -* `polynomial.is_eisenstein_at f 𝓟`: the property of being Eisenstein at `𝓟`. +In this file we gather more miscellaneous results about Eisenstein polynomials ## Main results -* `polynomial.is_eisenstein_at.irreducible`: if a primitive `f` satisfies `f.is_eisenstein_at 𝓟`, - where `𝓟.is_prime`, then `f` is irreducible. * `mem_adjoin_of_smul_prime_pow_smul_of_minpoly_is_eiseinstein_at`: let `K` be the field of fraction of an integrally closed domain `R` and let `L` be a separable extension of `K`, generated by an integral power basis `B` such that the minimal polynomial of `B.gen` is Eisenstein at `p`. Given @@ -28,11 +20,6 @@ Given an ideal `𝓟` of a commutative semiring `R`, we say that a polynomial `f Together with `algebra.discr_mul_is_integral_mem_adjoin` this result often allows to compute the ring of integers of `L`. -## Implementation details -We also define a notion `is_weakly_eisenstein_at` requiring only that -`∀ n < f.nat_degree → f.coeff n ∈ 𝓟`. This makes certain results slightly more general and it is -useful since it is sometimes better behaved (for example it is stable under `polynomial.map`). - -/ universes u v w z @@ -43,203 +30,6 @@ open ideal algebra finset open_locale big_operators polynomial -namespace polynomial - -/-- Given an ideal `𝓟` of a commutative semiring `R`, we say that a polynomial `f : R[X]` -is *weakly Eisenstein at `𝓟`* if `∀ n, n < f.nat_degree → f.coeff n ∈ 𝓟`. -/ -@[mk_iff] structure is_weakly_eisenstein_at [comm_semiring R] (f : R[X]) (𝓟 : ideal R) : - Prop := (mem : ∀ {n}, n < f.nat_degree → f.coeff n ∈ 𝓟) - -/-- Given an ideal `𝓟` of a commutative semiring `R`, we say that a polynomial `f : R[X]` -is *Eisenstein at `𝓟`* if `f.leading_coeff ∉ 𝓟`, `∀ n, n < f.nat_degree → f.coeff n ∈ 𝓟` and -`f.coeff 0 ∉ 𝓟 ^ 2`. -/ -@[mk_iff] structure is_eisenstein_at [comm_semiring R] (f : R[X]) (𝓟 : ideal R) : Prop := -(leading : f.leading_coeff ∉ 𝓟) -(mem : ∀ {n}, n < f.nat_degree → f.coeff n ∈ 𝓟) -(not_mem : f.coeff 0 ∉ 𝓟 ^ 2) - -namespace is_weakly_eisenstein_at - -section comm_semiring - -variables [comm_semiring R] {𝓟 : ideal R} {f : R[X]} (hf : f.is_weakly_eisenstein_at 𝓟) - -include hf - -lemma map {A : Type v} [comm_ring A] (φ : R →+* A) : (f.map φ).is_weakly_eisenstein_at (𝓟.map φ) := -begin - refine (is_weakly_eisenstein_at_iff _ _).2 (λ n hn, _), - rw [coeff_map], - exact mem_map_of_mem _ (hf.mem (lt_of_lt_of_le hn (nat_degree_map_le _ _))) -end - -end comm_semiring - -section comm_ring - -variables [comm_ring R] {𝓟 : ideal R} {f : R[X]} (hf : f.is_weakly_eisenstein_at 𝓟) -variables {S : Type v} [comm_ring S] [algebra R S] - -section principal - -variable {p : R} - -local notation `P` := submodule.span R {p} - -lemma exists_mem_adjoin_mul_eq_pow_nat_degree {x : S} (hx : aeval x f = 0) - (hmo : f.monic) (hf : f.is_weakly_eisenstein_at P) : ∃ y ∈ adjoin R ({x} : set S), - (algebra_map R S) p * y = x ^ (f.map (algebra_map R S)).nat_degree := -begin - rw [aeval_def, polynomial.eval₂_eq_eval_map, eval_eq_sum_range, range_add_one, - sum_insert not_mem_range_self, sum_range, (hmo.map - (algebra_map R S)).coeff_nat_degree, one_mul] at hx, - replace hx := eq_neg_of_add_eq_zero_left hx, - have : ∀ n < f.nat_degree, p ∣ f.coeff n, - { intros n hn, - refine mem_span_singleton.1 (by simpa using hf.mem hn) }, - choose! φ hφ using this, - conv_rhs at hx { congr, congr, skip, funext, - rw [fin.coe_eq_val, coeff_map, hφ i.1 (lt_of_lt_of_le i.2 (nat_degree_map_le _ _)), - ring_hom.map_mul, mul_assoc] }, - rw [hx, ← mul_sum, neg_eq_neg_one_mul, ← mul_assoc (-1 : S), mul_comm (-1 : S), mul_assoc], - refine ⟨-1 * ∑ (i : fin (f.map (algebra_map R S)).nat_degree), - (algebra_map R S) (φ i.1) * x ^ i.1, _, rfl⟩, - exact subalgebra.mul_mem _ (subalgebra.neg_mem _ (subalgebra.one_mem _)) - (subalgebra.sum_mem _ (λ i hi, subalgebra.mul_mem _ (subalgebra.algebra_map_mem _ _) - (subalgebra.pow_mem _ (subset_adjoin (set.mem_singleton x)) _))) -end - -lemma exists_mem_adjoin_mul_eq_pow_nat_degree_le {x : S} (hx : aeval x f = 0) - (hmo : f.monic) (hf : f.is_weakly_eisenstein_at P) : - ∀ i, (f.map (algebra_map R S)).nat_degree ≤ i → - ∃ y ∈ adjoin R ({x} : set S), (algebra_map R S) p * y = x ^ i := -begin - intros i hi, - obtain ⟨k, hk⟩ := exists_add_of_le hi, - rw [hk, pow_add], - obtain ⟨y, hy, H⟩ := exists_mem_adjoin_mul_eq_pow_nat_degree hx hmo hf, - refine ⟨y * x ^ k, _, _⟩, - { exact subalgebra.mul_mem _ hy (subalgebra.pow_mem _ (subset_adjoin (set.mem_singleton x)) _) }, - { rw [← mul_assoc _ y, H] } -end - -end principal - -include hf - -lemma pow_nat_degree_le_of_root_of_monic_mem {x : R} (hroot : is_root f x) (hmo : f.monic) : - ∀ i, f.nat_degree ≤ i → x ^ i ∈ 𝓟 := -begin - intros i hi, - obtain ⟨k, hk⟩ := exists_add_of_le hi, - rw [hk, pow_add], - suffices : x ^ f.nat_degree ∈ 𝓟, - { exact mul_mem_right (x ^ k) 𝓟 this }, - rw [is_root.def, eval_eq_sum_range, finset.range_add_one, finset.sum_insert - finset.not_mem_range_self, finset.sum_range, hmo.coeff_nat_degree, one_mul] at hroot, - rw [eq_neg_of_add_eq_zero_left hroot, neg_mem_iff], - refine submodule.sum_mem _ (λ i hi, mul_mem_right _ _ (hf.mem (fin.is_lt i))) -end - -lemma pow_nat_degree_le_of_aeval_zero_of_monic_mem_map {x : S} (hx : aeval x f = 0) - (hmo : f.monic) : - ∀ i, (f.map (algebra_map R S)).nat_degree ≤ i → x ^ i ∈ 𝓟.map (algebra_map R S) := -begin - suffices : x ^ (f.map (algebra_map R S)).nat_degree ∈ 𝓟.map (algebra_map R S), - { intros i hi, - obtain ⟨k, hk⟩ := exists_add_of_le hi, - rw [hk, pow_add], - refine mul_mem_right _ _ this }, - rw [aeval_def, eval₂_eq_eval_map, ← is_root.def] at hx, - refine pow_nat_degree_le_of_root_of_monic_mem (hf.map _) hx (hmo.map _) _ rfl.le -end - -end comm_ring - -end is_weakly_eisenstein_at - -section scale_roots - -variables {A : Type*} [comm_ring R] [comm_ring A] - -lemma scale_roots.is_weakly_eisenstein_at (p : R[X]) {x : R} - {P : ideal R} (hP : x ∈ P) : (scale_roots p x).is_weakly_eisenstein_at P := -begin - refine ⟨λ i hi, _⟩, - rw coeff_scale_roots, - rw [nat_degree_scale_roots, ← tsub_pos_iff_lt] at hi, - exact ideal.mul_mem_left _ _ (ideal.pow_mem_of_mem P hP _ hi) -end - -lemma dvd_pow_nat_degree_of_eval₂_eq_zero {f : R →+* A} - (hf : function.injective f) {p : R[X]} (hp : p.monic) (x y : R) (z : A) - (h : p.eval₂ f z = 0) (hz : f x * z = f y) : x ∣ y ^ p.nat_degree := -begin - rw [← nat_degree_scale_roots p x, ← ideal.mem_span_singleton], - refine (scale_roots.is_weakly_eisenstein_at _ (ideal.mem_span_singleton.mpr $ dvd_refl x)) - .pow_nat_degree_le_of_root_of_monic_mem _ ((monic_scale_roots_iff x).mpr hp) _ le_rfl, - rw injective_iff_map_eq_zero' at hf, - have := scale_roots_eval₂_eq_zero f h, - rwa [hz, polynomial.eval₂_at_apply, hf] at this -end - -lemma dvd_pow_nat_degree_of_aeval_eq_zero [algebra R A] [nontrivial A] - [no_zero_smul_divisors R A] {p : R[X]} (hp : p.monic) (x y : R) (z : A) - (h : polynomial.aeval z p = 0) (hz : z * algebra_map R A x = algebra_map R A y) : - x ∣ y ^ p.nat_degree := -dvd_pow_nat_degree_of_eval₂_eq_zero (no_zero_smul_divisors.algebra_map_injective R A) - hp x y z h ((mul_comm _ _).trans hz) - -end scale_roots - -namespace is_eisenstein_at - -section comm_semiring - -variables [comm_semiring R] {𝓟 : ideal R} {f : R[X]} (hf : f.is_eisenstein_at 𝓟) - -lemma _root_.polynomial.monic.leading_coeff_not_mem (hf : f.monic) (h : 𝓟 ≠ ⊤) : - ¬f.leading_coeff ∈ 𝓟 := -hf.leading_coeff.symm ▸ (ideal.ne_top_iff_one _).1 h - -lemma _root_.polynomial.monic.is_eisenstein_at_of_mem_of_not_mem (hf : f.monic) (h : 𝓟 ≠ ⊤) - (hmem : ∀ {n}, n < f.nat_degree → f.coeff n ∈ 𝓟) (hnot_mem : f.coeff 0 ∉ 𝓟 ^ 2) : - f.is_eisenstein_at 𝓟 := -{ leading := hf.leading_coeff_not_mem h, - mem := λ n hn, hmem hn, - not_mem := hnot_mem } - -include hf - -lemma is_weakly_eisenstein_at : is_weakly_eisenstein_at f 𝓟 := ⟨λ _, hf.mem⟩ - -lemma coeff_mem {n : ℕ} (hn : n ≠ f.nat_degree) : f.coeff n ∈ 𝓟 := -begin - cases ne_iff_lt_or_gt.1 hn, - { exact hf.mem h }, - { rw [coeff_eq_zero_of_nat_degree_lt h], - exact ideal.zero_mem _} -end - -end comm_semiring - -section is_domain - -variables [comm_ring R] [is_domain R] {𝓟 : ideal R} {f : R[X]} (hf : f.is_eisenstein_at 𝓟) - -/-- If a primitive `f` satisfies `f.is_eisenstein_at 𝓟`, where `𝓟.is_prime`, then `f` is -irreducible. -/ -lemma irreducible (hprime : 𝓟.is_prime) (hu : f.is_primitive) - (hfd0 : 0 < f.nat_degree) : irreducible f := -irreducible_of_eisenstein_criterion hprime hf.leading (λ n hn, hf.mem (coe_lt_degree.1 hn)) - (nat_degree_pos_iff_degree_pos.1 hfd0) hf.not_mem hu - -end is_domain - -end is_eisenstein_at - -end polynomial - section cyclotomic variables (p : ℕ) @@ -329,7 +119,7 @@ section is_integral variables {K : Type v} {L : Type z} {p : R} [comm_ring R] [field K] [field L] variables [algebra K L] [algebra R L] [algebra R K] [is_scalar_tower R K L] [is_separable K L] -variables [is_domain R] [normalized_gcd_monoid R] [is_fraction_ring R K] [is_integrally_closed R] +variables [is_domain R] [is_fraction_ring R K] [is_integrally_closed R] local notation `𝓟` := submodule.span R {p} @@ -351,7 +141,7 @@ begin have finrank_K_L : finite_dimensional.finrank K L = B.dim := B.finrank, have deg_K_P : (minpoly K B.gen).nat_degree = B.dim := B.nat_degree_minpoly, have deg_R_P : P.nat_degree = B.dim, - { rw [← deg_K_P, minpoly.gcd_domain_eq_field_fractions' K hBint, + { rw [← deg_K_P, minpoly.is_integrally_closed_eq_field_fractions' K hBint, (minpoly.monic hBint).nat_degree_map (algebra_map R K)] }, choose! f hf using hei.is_weakly_eisenstein_at.exists_mem_adjoin_mul_eq_pow_nat_degree_le (minpoly.aeval R B.gen) (minpoly.monic hBint), @@ -393,7 +183,7 @@ begin ... = _ : _, { simp only [algebra.smul_def, algebra_map_apply R K L, algebra.norm_algebra_map, _root_.map_mul, _root_.map_pow, finrank_K_L, power_basis.norm_gen_eq_coeff_zero_minpoly, - minpoly.gcd_domain_eq_field_fractions' K hBint, coeff_map, ← hn], + minpoly.is_integrally_closed_eq_field_fractions' K hBint, coeff_map, ← hn], ring_exp }, swap, { simp_rw [← smul_sum, ← smul_sub, algebra.smul_def p, algebra_map_apply R K L, _root_.map_mul, algebra.norm_algebra_map, finrank_K_L, hr, ← hn] }, @@ -527,8 +317,8 @@ begin rw [nat.succ_eq_add_one, add_assoc, ← nat.add_sub_assoc H, ← add_assoc, add_comm (j + 1), nat.add_sub_add_left, ← nat.add_sub_assoc, nat.add_sub_add_left, hP, ← (minpoly.monic hBint).nat_degree_map (algebra_map R K), - ← minpoly.gcd_domain_eq_field_fractions' K hBint, nat_degree_minpoly, hn, nat.sub_one, - nat.pred_succ], + ← minpoly.is_integrally_closed_eq_field_fractions' K hBint, nat_degree_minpoly, hn, + nat.sub_one, nat.pred_succ], linarith }, -- Using `hQ : aeval B.gen Q = p • z`, we write `p • z` as a sum of terms of degree less than @@ -564,8 +354,8 @@ begin obtain ⟨r, hr⟩ := is_integral_iff.1 (is_integral_norm K hintsum), rw [algebra.smul_def, mul_assoc, ← mul_sub, _root_.map_mul, algebra_map_apply R K L, map_pow, algebra.norm_algebra_map, _root_.map_mul, algebra_map_apply R K L, algebra.norm_algebra_map, - finrank B, ← hr, - power_basis.norm_gen_eq_coeff_zero_minpoly, minpoly.gcd_domain_eq_field_fractions' K hBint, + finrank B, ← hr, power_basis.norm_gen_eq_coeff_zero_minpoly, + minpoly.is_integrally_closed_eq_field_fractions' K hBint, coeff_map, show (-1 : K) = algebra_map R K (-1), by simp, ← map_pow, ← map_pow, ← _root_.map_mul, ← map_pow, ← _root_.map_mul, ← map_pow, ← _root_.map_mul] at hQ, diff --git a/src/ring_theory/roots_of_unity.lean b/src/ring_theory/roots_of_unity.lean index d4ba7b344f82f..c21e28e689fb0 100644 --- a/src/ring_theory/roots_of_unity.lean +++ b/src/ring_theory/roots_of_unity.lean @@ -6,9 +6,10 @@ Authors: Johan Commelin import algebra.char_p.two import algebra.ne_zero +import algebra.gcd_monoid.integrally_closed import data.polynomial.ring_division import field_theory.finite.basic -import field_theory.minpoly.gcd_monoid +import field_theory.minpoly.is_integrally_closed import field_theory.separable import group_theory.specific_groups.cyclic import number_theory.divisors @@ -908,7 +909,8 @@ lemma minpoly_dvd_X_pow_sub_one : minpoly ℤ μ ∣ X ^ n - 1 := begin rcases n.eq_zero_or_pos with rfl | hpos, { simp }, - apply minpoly.gcd_domain_dvd (is_integral h hpos) (monic_X_pow_sub_C 1 hpos.ne').ne_zero, + letI : is_integrally_closed ℤ := gcd_monoid.to_is_integrally_closed, + apply minpoly.is_integrally_closed_dvd (is_integral h hpos), simp only [((is_primitive_root.iff_def μ n).mp h).left, aeval_X_pow, eq_int_cast, int.cast_one, aeval_one, alg_hom.map_sub, sub_self] end @@ -933,17 +935,13 @@ lemma squarefree_minpoly_mod {p : ℕ} [fact p.prime] (hdiv : ¬ p ∣ n) : (separable_minpoly_mod h hdiv).squarefree /- Let `P` be the minimal polynomial of a root of unity `μ` and `Q` be the minimal polynomial of -`μ ^ p`, where `p` is a prime that does not divide `n`. Then `P` divides `expand ℤ p Q`. -/ -lemma minpoly_dvd_expand {p : ℕ} (hprime : nat.prime p) (hdiv : ¬ p ∣ n) : - minpoly ℤ μ ∣ expand ℤ p (minpoly ℤ (μ ^ p)) := +`μ ^ p`, where `p` is a natural number that does not divide `n`. Then `P` divides `expand ℤ p Q`. -/ +lemma minpoly_dvd_expand {p : ℕ} (hdiv : ¬ p ∣ n) : minpoly ℤ μ ∣ expand ℤ p (minpoly ℤ (μ ^ p)) := begin rcases n.eq_zero_or_pos with rfl | hpos, { simp * at *, }, - refine minpoly.gcd_domain_dvd (h.is_integral hpos) _ _, - { apply monic.ne_zero, - rw [polynomial.monic, leading_coeff, nat_degree_expand, mul_comm, coeff_expand_mul' - (nat.prime.pos hprime), ← leading_coeff, ← polynomial.monic], - exact minpoly.monic (is_integral (pow_of_prime h hprime hdiv) hpos) }, + letI : is_integrally_closed ℤ := gcd_monoid.to_is_integrally_closed, + refine minpoly.is_integrally_closed_dvd (h.is_integral hpos) _, { rw [aeval_def, coe_expand, ← comp, eval₂_eq_eval_map, map_comp, polynomial.map_pow, map_X, eval_comp, eval_pow, eval_X, ← eval₂_eq_eval_map, ← aeval_def], exact minpoly.aeval _ _ } @@ -961,7 +959,7 @@ begin by rw [← zmod.expand_card, map_expand], rw [hfrob], apply ring_hom.map_dvd (map_ring_hom (int.cast_ring_hom (zmod p))), - exact minpoly_dvd_expand h hprime.1 hdiv + exact minpoly_dvd_expand h hdiv end /- Let `P` be the minimal polynomial of a root of unity `μ` and `Q` be the minimal polynomial of From 959c3b69db8a8b404d5813421f2e6ca8660d19e5 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Mon, 23 Jan 2023 10:09:00 +0000 Subject: [PATCH 0328/1291] doc(computability/primrec): remove doc nolints (#16018) We just add some documentation. --- src/computability/primrec.lean | 10 ++++++++-- 1 file changed, 8 insertions(+), 2 deletions(-) diff --git a/src/computability/primrec.lean b/src/computability/primrec.lean index e6b4673132d2f..a6e4b7a34941f 100644 --- a/src/computability/primrec.lean +++ b/src/computability/primrec.lean @@ -31,17 +31,20 @@ open denumerable encodable function namespace nat +/-- The non-dependent recursor on naturals. -/ def elim {C : Sort*} : C → (ℕ → C → C) → ℕ → C := @nat.rec (λ _, C) @[simp] theorem elim_zero {C} (a f) : @nat.elim C a f 0 = a := rfl @[simp] theorem elim_succ {C} (a f n) : @nat.elim C a f (succ n) = f n (nat.elim a f n) := rfl +/-- Cases on whether the input is 0 or a successor. -/ def cases {C : Sort*} (a : C) (f : ℕ → C) : ℕ → C := nat.elim a (λ n _, f n) @[simp] theorem cases_zero {C} (a f) : @nat.cases C a f 0 = a := rfl @[simp] theorem cases_succ {C} (a f n) : @nat.cases C a f (succ n) = f n := rfl +/-- Calls the given function on a pair of entries `n`, encoded via the pairing function. -/ @[simp, reducible] def unpaired {α} (f : ℕ → ℕ → α) (n : ℕ) : α := f n.unpair.1 n.unpair.2 @@ -72,7 +75,7 @@ theorem prec1 {f} (m : ℕ) (hf : primrec f) : primrec (λ n, n.elim m (λ y IH, f $ mkpair y IH)) := ((prec (const m) (hf.comp right)).comp (zero.pair primrec.id)).of_eq $ -λ n, by simp; dsimp; rw [unpair_mkpair] +λ n, by simp theorem cases1 {f} (m : ℕ) (hf : primrec f) : primrec (nat.cases m f) := (prec1 m (hf.comp left)).of_eq $ by simp [cases] @@ -121,6 +124,7 @@ open nat.primrec @[priority 10] instance of_denumerable (α) [denumerable α] : primcodable α := ⟨succ.of_eq $ by simp⟩ +/-- Builds a `primcodable` instance from an equivalence to a `primcodable` type. -/ def of_equiv (α) {β} [primcodable α] (e : β ≃ α) : primcodable β := { prim := (primcodable.prim α).of_eq $ λ n, show encode (decode α n) = @@ -1018,6 +1022,7 @@ variables {α : Type*} {β : Type*} variables [primcodable α] [primcodable β] open primrec +/-- A subtype of a primitive recursive predicate is `primcodable`. -/ def subtype {p : α → Prop} [decidable_pred p] (hp : primrec_pred p) : primcodable (subtype p) := ⟨have primrec (λ n, (decode α n).bind (λ a, option.guard p a)), @@ -1231,7 +1236,8 @@ theorem tail {n f} (hf : @primrec' n f) : @primrec' n.succ (λ v, f v.tail) := (hf.comp _ (λ i, @nth _ i.succ)).of_eq $ λ v, by rw [← of_fn_nth v.tail]; congr; funext i; simp -def vec {n m} (f : vector ℕ n → vector ℕ m) := +/-- A function from vectors to vectors is primitive recursive when all of its projections are. -/ +def vec {n m} (f : vector ℕ n → vector ℕ m) : Prop := ∀ i, primrec' (λ v, (f v).nth i) protected theorem nil {n} : @vec n 0 (λ _, nil) := λ i, i.elim0 From ed989ff568099019c6533a4d94b27d852a5710d8 Mon Sep 17 00:00:00 2001 From: prakol16 Date: Mon, 23 Jan 2023 10:09:01 +0000 Subject: [PATCH 0329/1291] feat(data/tree): Add definitions about binary trees; special support for `tree unit` (#16715) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Provide convenience API for `tree unit`. In particular, define a recursor for `tree unit` as well as other helper lemmas. In addition, define common functions on binary trees such as `nodes`, `num_leaves`, `height`, and prove basic lemmas about these functions, such as the fact that the number of leaves is one more than the number of internal nodes and the height is at most the number of nodes. Co-authored-by: Yaël Dillies --- src/data/tree.lean | 57 ++++++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 57 insertions(+) diff --git a/src/data/tree.lean b/src/data/tree.lean index 6e3e388016a30..50c55cd712a64 100644 --- a/src/data/tree.lean +++ b/src/data/tree.lean @@ -5,6 +5,7 @@ Authors: Mario Carneiro, Wojciech Nawrocki -/ import data.rbtree.init import data.num.basic +import order.basic /-! # Binary tree @@ -13,6 +14,10 @@ Provides binary tree storage for values of any type, with O(lg n) retrieval. See also `data.rbtree` for red-black trees - this version allows more operations to be defined and is better suited for in-kernel computation. +We also specialize for `tree unit`, which is a binary tree without any +additional data. We provide the notation `a △ b` for making a `tree unit` with children +`a` and `b`. + ## References @@ -81,4 +86,56 @@ def map {β} (f : α → β) : tree α → tree β | nil := nil | (node a l r) := node (f a) (map l) (map r) +/-- The number of internal nodes (i.e. not including leaves) of a binary tree -/ +@[simp] def num_nodes : tree α → ℕ +| nil := 0 +| (node _ a b) := a.num_nodes + b.num_nodes + 1 + +/-- The number of leaves of a binary tree -/ +@[simp] def num_leaves : tree α → ℕ +| nil := 1 +| (node _ a b) := a.num_leaves + b.num_leaves + +/-- The height - length of the longest path from the root - of a binary tree -/ +@[simp] def height : tree α → ℕ +| nil := 0 +| (node _ a b) := max a.height b.height + 1 + +lemma num_leaves_eq_num_nodes_succ (x : tree α) : x.num_leaves = x.num_nodes + 1 := +by { induction x; simp [*, nat.add_comm, nat.add_assoc, nat.add_left_comm], } + +lemma num_leaves_pos (x : tree α) : 0 < x.num_leaves := +by { rw num_leaves_eq_num_nodes_succ, exact x.num_nodes.zero_lt_succ, } + +lemma height_le_num_nodes : ∀ (x : tree α), x.height ≤ x.num_nodes +| nil := le_rfl +| (node _ a b) := nat.succ_le_succ + (max_le + (trans a.height_le_num_nodes $ a.num_nodes.le_add_right _) + (trans b.height_le_num_nodes $ b.num_nodes.le_add_left _)) + +/-- The left child of the tree, or `nil` if the tree is `nil` -/ +@[simp] def left : tree α → tree α +| nil := nil +| (node _ l r) := l + +/-- The right child of the tree, or `nil` if the tree is `nil` -/ +@[simp] def right : tree α → tree α +| nil := nil +| (node _ l r) := r + +/- Notation for making a node with `unit` data -/ +localized "infixr ` △ `:65 := tree.node ()" in tree + +/-- Recursion on `tree unit`; allows for a better `induction` which does not have to worry + about the element of type `α = unit` -/ +@[elab_as_eliminator] +def unit_rec_on {motive : tree unit → Sort*} (t : tree unit) (base : motive nil) + (ind : ∀ x y, motive x → motive y → motive (x △ y)) : motive t := +t.rec_on base (λ u, u.rec_on (by exact ind)) + +lemma left_node_right_eq_self : ∀ {x : tree unit} (hx : x ≠ nil), x.left △ x.right = x +| nil h := by trivial +| (a △ b) _ := rfl + end tree From 6e492fdbed778ea815798a5a0c6cb83df8486265 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 23 Jan 2023 10:09:02 +0000 Subject: [PATCH 0330/1291] fix(*/finsupp/*): add missing decidable arguments in lemma statements (#18251) I missed these files in #18241 --- src/algebra/big_operators/finsupp.lean | 35 ++++++++++++++------ src/algebra/category/Module/adjunctions.lean | 10 +++--- src/algebra/monoid_algebra/basic.lean | 24 ++++++++------ src/data/finsupp/basic.lean | 19 ++++++++--- src/data/finsupp/indicator.lean | 17 ++++++---- src/data/finsupp/interval.lean | 10 ++++-- src/data/finsupp/order.lean | 9 +++-- src/data/finsupp/to_dfinsupp.lean | 13 +++++--- src/data/mv_polynomial/derivation.lean | 2 +- src/ring_theory/ideal/operations.lean | 2 +- 10 files changed, 93 insertions(+), 48 deletions(-) diff --git a/src/algebra/big_operators/finsupp.lean b/src/algebra/big_operators/finsupp.lean index a4b71c17979e1..ded7e78f17f9f 100644 --- a/src/algebra/big_operators/finsupp.lean +++ b/src/algebra/big_operators/finsupp.lean @@ -19,7 +19,7 @@ This file contains theorems relevant to big operators in finitely supported func noncomputable theory open finset function -open_locale classical big_operators +open_locale big_operators variables {α ι γ A B C : Type*} [add_comm_monoid A] [add_comm_monoid B] [add_comm_monoid C] variables {t : ι → A → C} (h0 : ∀ i, t i 0 = 0) (h1 : ∀ i x y, t i (x + y) = t i x + t i y) @@ -86,7 +86,11 @@ by { dsimp [finsupp.prod], rw f.support.prod_ite_eq, } @[simp] lemma sum_ite_self_eq [decidable_eq α] {N : Type*} [add_comm_monoid N] (f : α →₀ N) (a : α) : f.sum (λ x v, ite (a = x) v 0) = f a := -by { convert f.sum_ite_eq a (λ x, id), simp [ite_eq_right_iff.2 eq.symm] } +begin + classical, + convert f.sum_ite_eq a (λ x, id), + simp [ite_eq_right_iff.2 eq.symm] +end /-- A restatement of `prod_ite_eq` with the equality test reversed. -/ @[simp, to_additive "A restatement of `sum_ite_eq` with the equality test reversed."] @@ -97,7 +101,11 @@ by { dsimp [finsupp.prod], rw f.support.prod_ite_eq', } @[simp] lemma sum_ite_self_eq' [decidable_eq α] {N : Type*} [add_comm_monoid N] (f : α →₀ N) (a : α) : f.sum (λ x v, ite (x = a) v 0) = f a := -by { convert f.sum_ite_eq' a (λ x, id), simp [ite_eq_right_iff.2 eq.symm] } +begin + classical, + convert f.sum_ite_eq' a (λ x, id), + simp [ite_eq_right_iff.2 eq.symm] +end @[simp] lemma prod_pow [fintype α] (f : α →₀ ℕ) (g : α → N) : f.prod (λ a b, g a ^ b) = ∏ a, g a ^ (f a) := @@ -121,6 +129,7 @@ single element `y ∈ f.support` to the sum over `erase y f`. "-/] lemma mul_prod_erase (f : α →₀ M) (y : α) (g : α → M → N) (hyf : y ∈ f.support) : g y (f y) * (erase y f).prod g = f.prod g := begin + classical, rw [finsupp.prod, finsupp.prod, ←finset.mul_prod_erase _ _ hyf, finsupp.support_erase, finset.prod_congr rfl], intros h hx, @@ -288,7 +297,7 @@ This is a more general version of `finsupp.prod_add_index'`; the latter has simp @[to_additive "Taking the product under `h` is an additive homomorphism of finsupps, if `h` is an additive homomorphism on the support. This is a more general version of `finsupp.sum_add_index'`; the latter has simpler hypotheses."] -lemma prod_add_index [add_zero_class M] [comm_monoid N] {f g : α →₀ M} +lemma prod_add_index [decidable_eq α] [add_zero_class M] [comm_monoid N] {f g : α →₀ M} {h : α → M → N} (h_zero : ∀ a ∈ f.support ∪ g.support, h a 0 = 1) (h_add : ∀ (a ∈ f.support ∪ g.support) b₁ b₂, h a (b₁ + b₂) = h a b₁ * h a b₂) : (f + g).prod h = f.prod h * g.prod h := @@ -309,7 +318,7 @@ This is a more specific version of `finsupp.sum_add_index` with simpler hypothes lemma prod_add_index' [add_zero_class M] [comm_monoid N] {f g : α →₀ M} {h : α → M → N} (h_zero : ∀a, h a 0 = 1) (h_add : ∀a b₁ b₂, h a (b₁ + b₂) = h a b₁ * h a b₂) : (f + g).prod h = f.prod h * g.prod h := -prod_add_index (λ a ha, h_zero a) (λ a ha, h_add a) +by classical; exact prod_add_index (λ a ha, h_zero a) (λ a ha, h_add a) @[simp] lemma sum_hom_add_index [add_zero_class M] [add_comm_monoid N] {f g : α →₀ M} (h : α → M →+ N) : @@ -401,8 +410,8 @@ lemma prod_finset_sum_index [add_comm_monoid M] [comm_monoid N] {s : finset ι} {g : ι → α →₀ M} {h : α → M → N} (h_zero : ∀a, h a 0 = 1) (h_add : ∀a b₁ b₂, h a (b₁ + b₂) = h a b₁ * h a b₂) : ∏ i in s, (g i).prod h = (∑ i in s, g i).prod h := -finset.induction_on s rfl $ λ a s has ih, -by rw [prod_insert has, ih, sum_insert has, prod_add_index' h_zero h_add] +finset.cons_induction_on s rfl $ λ a s has ih, +by rw [prod_cons, ih, sum_cons, prod_add_index' h_zero h_add] @[to_additive] lemma prod_sum_index @@ -420,10 +429,12 @@ lemma multiset_sum_sum_index multiset.induction_on f rfl $ assume a s ih, by rw [multiset.sum_cons, multiset.map_cons, multiset.sum_cons, sum_add_index' h₀ h₁, ih] -lemma support_sum_eq_bUnion {α : Type*} {ι : Type*} {M : Type*} [add_comm_monoid M] +lemma support_sum_eq_bUnion {α : Type*} {ι : Type*} {M : Type*} [decidable_eq α] + [add_comm_monoid M] {g : ι → α →₀ M} (s : finset ι) (h : ∀ i₁ i₂, i₁ ≠ i₂ → disjoint (g i₁).support (g i₂).support) : (∑ i in s, g i).support = s.bUnion (λ i, (g i).support) := begin + classical, apply finset.induction_on s, { simp }, { intros i s hi, @@ -455,14 +466,18 @@ have ∀ {f1 f2 : α →₀ M}, disjoint f1.support f2.support → ∏ x in f1.support, g x (f1 x + f2 x) = f1.prod g := λ f1 f2 hd, finset.prod_congr rfl (λ x hx, by simp only [not_mem_support_iff.mp (disjoint_left.mp hd hx), add_zero]), -by simp_rw [← this hd, ← this hd.symm, - add_comm (f2 _), finsupp.prod, support_add_eq hd, prod_union hd, add_apply] +begin + classical, + simp_rw [← this hd, ← this hd.symm, + add_comm (f2 _), finsupp.prod, support_add_eq hd, prod_union hd, add_apply] +end lemma prod_dvd_prod_of_subset_of_dvd [add_comm_monoid M] [comm_monoid N] {f1 f2 : α →₀ M} {g1 g2 : α → M → N} (h1 : f1.support ⊆ f2.support) (h2 : ∀ (a : α), a ∈ f1.support → g1 a (f1 a) ∣ g2 a (f2 a)) : f1.prod g1 ∣ f2.prod g2 := begin + classical, simp only [finsupp.prod, finsupp.prod_mul], rw [←sdiff_union_of_subset h1, prod_union sdiff_disjoint], apply dvd_mul_of_dvd_right, diff --git a/src/algebra/category/Module/adjunctions.lean b/src/algebra/category/Module/adjunctions.lean index e3d6d7980eb0e..e03b0271ee9ce 100644 --- a/src/algebra/category/Module/adjunctions.lean +++ b/src/algebra/category/Module/adjunctions.lean @@ -196,14 +196,14 @@ instance : preadditive (Free R C) := { hom_group := λ X Y, finsupp.add_comm_group, add_comp' := λ X Y Z f f' g, begin dsimp, - rw [finsupp.sum_add_index]; + rw [finsupp.sum_add_index']; { simp [add_mul], } end, comp_add' := λ X Y Z f g g', begin dsimp, rw ← finsupp.sum_add, congr, ext r h, - rw [finsupp.sum_add_index]; + rw [finsupp.sum_add_index']; { simp [mul_add], }, end, } @@ -257,7 +257,7 @@ def lift (F : C ⥤ D) : Free R C ⥤ D := { simp only [limits.zero_comp, sum_zero_index] }, { intros f₁ f₂ w₁ w₂, rw add_comp, - rw [finsupp.sum_add_index, finsupp.sum_add_index], + rw [finsupp.sum_add_index', finsupp.sum_add_index'], { simp only [w₁, w₂, add_comp] }, { intros, rw zero_smul }, { intros, simp only [add_smul], }, @@ -268,7 +268,7 @@ def lift (F : C ⥤ D) : Free R C ⥤ D := { simp only [limits.comp_zero, sum_zero_index] }, { intros f₁ f₂ w₁ w₂, rw comp_add, - rw [finsupp.sum_add_index, finsupp.sum_add_index], + rw [finsupp.sum_add_index', finsupp.sum_add_index'], { simp only [w₁, w₂, comp_add], }, { intros, rw zero_smul }, { intros, simp only [add_smul], }, @@ -287,7 +287,7 @@ by simp instance lift_additive (F : C ⥤ D) : (lift R F).additive := { map_add' := λ X Y f g, begin dsimp, - rw finsupp.sum_add_index; simp [add_smul] + rw finsupp.sum_add_index'; simp [add_smul] end, } instance lift_linear (F : C ⥤ D) : (lift R F).linear R := diff --git a/src/algebra/monoid_algebra/basic.lean b/src/algebra/monoid_algebra/basic.lean index 5eac4ef411e8b..a09ea8d935603 100644 --- a/src/algebra/monoid_algebra/basic.lean +++ b/src/algebra/monoid_algebra/basic.lean @@ -109,11 +109,13 @@ instance : non_unital_non_assoc_semiring (monoid_algebra k G) := { zero := 0, mul := (*), add := (+), - left_distrib := assume f g h, by simp only [mul_def, sum_add_index, mul_add, mul_zero, - single_zero, single_add, eq_self_iff_true, forall_true_iff, forall_3_true_iff, sum_add], - right_distrib := assume f g h, by simp only [mul_def, sum_add_index, add_mul, zero_mul, - single_zero, single_add, eq_self_iff_true, forall_true_iff, forall_3_true_iff, sum_zero, - sum_add], + left_distrib := assume f g h, by haveI := classical.dec_eq G; + simp only [mul_def, sum_add_index, mul_add, mul_zero, + single_zero, single_add, eq_self_iff_true, forall_true_iff, forall_3_true_iff, sum_add], + right_distrib := assume f g h, by haveI := classical.dec_eq G; + simp only [mul_def, sum_add_index, add_mul, zero_mul, + single_zero, single_add, eq_self_iff_true, forall_true_iff, forall_3_true_iff, sum_zero, + sum_add], zero_mul := assume f, by simp only [mul_def, sum_zero_index], mul_zero := assume f, by simp only [mul_def, sum_zero_index, sum_zero], .. finsupp.add_comm_monoid } @@ -924,11 +926,13 @@ instance : non_unital_non_assoc_semiring (add_monoid_algebra k G) := { zero := 0, mul := (*), add := (+), - left_distrib := assume f g h, by simp only [mul_def, sum_add_index, mul_add, mul_zero, - single_zero, single_add, eq_self_iff_true, forall_true_iff, forall_3_true_iff, sum_add], - right_distrib := assume f g h, by simp only [mul_def, sum_add_index, add_mul, mul_zero, zero_mul, - single_zero, single_add, eq_self_iff_true, forall_true_iff, forall_3_true_iff, sum_zero, - sum_add], + left_distrib := assume f g h, by haveI := classical.dec_eq G; + simp only [mul_def, sum_add_index, mul_add, mul_zero, + single_zero, single_add, eq_self_iff_true, forall_true_iff, forall_3_true_iff, sum_add], + right_distrib := assume f g h, by haveI := classical.dec_eq G; + simp only [mul_def, sum_add_index, add_mul, mul_zero, zero_mul, + single_zero, single_add, eq_self_iff_true, forall_true_iff, forall_3_true_iff, sum_zero, + sum_add], zero_mul := assume f, by simp only [mul_def, sum_zero_index], mul_zero := assume f, by simp only [mul_def, sum_zero_index, sum_zero], nsmul := λ n f, n • f, diff --git a/src/data/finsupp/basic.lean b/src/data/finsupp/basic.lean index c3ea2d6e43d42..f2cba85f09fa7 100644 --- a/src/data/finsupp/basic.lean +++ b/src/data/finsupp/basic.lean @@ -733,8 +733,9 @@ lemma prod_option_index [add_comm_monoid M] [comm_monoid N] (h_add : ∀ o m₁ m₂, b o (m₁ + m₂) = b o m₁ * b o m₂) : f.prod b = b none (f none) * f.some.prod (λ a, b (option.some a)) := begin + classical, apply induction_linear f, - { simp [h_zero], }, + { simp [some_zero, h_zero], }, { intros f₁ f₂ h₁ h₂, rw [finsupp.prod_add_index, h₁, h₂, some_add, finsupp.prod_add_index], simp only [h_add, pi.add_apply, finsupp.coe_add], @@ -1053,10 +1054,18 @@ f.sum $ λa g, g.sum $ λb c, single (a, b) c /-- `finsupp_prod_equiv` defines the `equiv` between `((α × β) →₀ M)` and `(α →₀ (β →₀ M))` given by currying and uncurrying. -/ def finsupp_prod_equiv : ((α × β) →₀ M) ≃ (α →₀ (β →₀ M)) := -by refine ⟨finsupp.curry, finsupp.uncurry, λ f, _, λ f, _⟩; simp only [ - finsupp.curry, finsupp.uncurry, sum_sum_index, sum_zero_index, sum_add_index, - sum_single_index, single_zero, single_add, eq_self_iff_true, forall_true_iff, - forall_3_true_iff, prod.mk.eta, (single_sum _ _ _).symm, sum_single] +{ to_fun := finsupp.curry, + inv_fun := finsupp.uncurry, + left_inv := λ f, begin + rw [finsupp.uncurry, sum_curry_index], + { simp_rw [prod.mk.eta, sum_single], }, + { intros, apply single_zero }, + { intros, apply single_add } + end, + right_inv := λ f, by simp only [ + finsupp.curry, finsupp.uncurry, sum_sum_index, sum_zero_index, sum_add_index, + sum_single_index, single_zero, single_add, eq_self_iff_true, forall_true_iff, + forall_3_true_iff, prod.mk.eta, (single_sum _ _ _).symm, sum_single] } lemma filter_curry (f : α × β →₀ M) (p : α → Prop) : (f.filter (λa:α×β, p a.1)).curry = f.curry.filter p := diff --git a/src/data/finsupp/indicator.lean b/src/data/finsupp/indicator.lean index d2fa77851a510..bf2a3f7ae1597 100644 --- a/src/data/finsupp/indicator.lean +++ b/src/data/finsupp/indicator.lean @@ -18,7 +18,6 @@ This file defines `finsupp.indicator` to help create finsupps from finsets. noncomputable theory open finset function -open_locale classical variables {ι α : Type*} @@ -27,20 +26,26 @@ variables [has_zero α] {s : finset ι} (f : Π i ∈ s, α) {i : ι} /-- Create an element of `ι →₀ α` from a finset `s` and a function `f` defined on this finset. -/ def indicator (s : finset ι) (f : Π i ∈ s, α) : ι →₀ α := -{ to_fun := λ i, if H : i ∈ s then f i H else 0, - support := (s.attach.filter $ λ i : s, f i.1 i.2 ≠ 0).map $ embedding.subtype _, +{ to_fun := λ i, by haveI := classical.dec_eq ι; exact + if H : i ∈ s then f i H else 0, + support := by haveI := classical.dec_eq α; exact + (s.attach.filter $ λ i : s, f i.1 i.2 ≠ 0).map (embedding.subtype _), mem_support_to_fun := λ i, begin + letI := classical.dec_eq α, rw [mem_map, dite_ne_right_iff], exact ⟨λ ⟨⟨j, hj⟩, hf, rfl⟩, ⟨hj, (mem_filter.1 hf).2⟩, λ ⟨hi, hf⟩, ⟨⟨i, hi⟩, mem_filter.2 $ ⟨mem_attach _ _, hf⟩, rfl⟩⟩, end } -lemma indicator_of_mem (hi : i ∈ s) (f : Π i ∈ s, α) : indicator s f i = f i hi := dif_pos hi -lemma indicator_of_not_mem (hi : i ∉ s) (f : Π i ∈ s, α) : indicator s f i = 0 := dif_neg hi +lemma indicator_of_mem (hi : i ∈ s) (f : Π i ∈ s, α) : indicator s f i = f i hi := +@dif_pos _ (id _) hi _ _ _ +lemma indicator_of_not_mem (hi : i ∉ s) (f : Π i ∈ s, α) : indicator s f i = 0 := +@dif_neg _ (id _) hi _ _ _ variables (s i) -@[simp] lemma indicator_apply : indicator s f i = if hi : i ∈ s then f i hi else 0 := rfl +@[simp] lemma indicator_apply [decidable_eq ι] : + indicator s f i = if hi : i ∈ s then f i hi else 0 := by convert rfl lemma indicator_injective : injective (λ f : Π i ∈ s, α, indicator s f) := begin diff --git a/src/data/finsupp/interval.lean b/src/data/finsupp/interval.lean index 99845f4cf77e3..985a3572b6e3a 100644 --- a/src/data/finsupp/interval.lean +++ b/src/data/finsupp/interval.lean @@ -53,7 +53,7 @@ variables [has_zero α] [partial_order α] [locally_finite_order α] {f g : ι /-- Pointwise `finset.Icc` bundled as a `finsupp`. -/ @[simps to_fun] def range_Icc (f g : ι →₀ α) : ι →₀ finset α := { to_fun := λ i, Icc (f i) (g i), - support := f.support ∪ g.support, + support := by haveI := classical.dec_eq ι; exact f.support ∪ g.support, mem_support_to_fun := λ i, begin rw [mem_union, ←not_iff_not, not_or_distrib, not_mem_support_iff, not_mem_support_iff, not_ne_iff], @@ -72,6 +72,7 @@ section partial_order variables [partial_order α] [has_zero α] [locally_finite_order α] (f g : ι →₀ α) instance : locally_finite_order (ι →₀ α) := +by haveI := classical.dec_eq ι; haveI := classical.dec_eq α; exact locally_finite_order.of_Icc (ι →₀ α) (λ f g, (f.support ∪ g.support).finsupp $ f.range_Icc g) (λ f g x, begin @@ -108,8 +109,11 @@ variables [canonically_ordered_add_monoid α] [locally_finite_order α] variables (f : ι →₀ α) lemma card_Iic : (Iic f).card = ∏ i in f.support, (Iic (f i)).card := -by simp_rw [Iic_eq_Icc, card_Icc, finsupp.bot_eq_zero, support_zero, empty_union, zero_apply, - bot_eq_zero] +begin + classical, + simp_rw [Iic_eq_Icc, card_Icc, finsupp.bot_eq_zero, support_zero, empty_union, zero_apply, + bot_eq_zero] +end lemma card_Iio : (Iio f).card = ∏ i in f.support, (Iic (f i)).card - 1 := by rw [card_Iio_eq_card_Iic_sub_one, card_Iic] diff --git a/src/data/finsupp/order.lean b/src/data/finsupp/order.lean index 2f404c709d442..5d6ba134d6fc8 100644 --- a/src/data/finsupp/order.lean +++ b/src/data/finsupp/order.lean @@ -19,7 +19,7 @@ This file lifts order structures on `α` to `ι →₀ α`. -/ noncomputable theory -open_locale classical big_operators +open_locale big_operators open finset @@ -116,7 +116,8 @@ by simp [ext_iff, forall_and_distrib] lemma le_iff' (f g : ι →₀ α) {s : finset ι} (hf : f.support ⊆ s) : f ≤ g ↔ ∀ i ∈ s, f i ≤ g i := ⟨λ h s hs, h s, -λ h s, if H : s ∈ f.support then h s (hf H) else (not_mem_support_iff.1 H).symm ▸ zero_le (g s)⟩ +λ h s, by classical; exact + if H : s ∈ f.support then h s (hf H) else (not_mem_support_iff.1 H).symm ▸ zero_le (g s)⟩ lemma le_iff (f g : ι →₀ α) : f ≤ g ↔ ∀ i ∈ f.support, f i ≤ g i := le_iff' f g $ subset.refl _ @@ -156,7 +157,8 @@ lemma support_tsub {f1 f2 : ι →₀ α} : (f1 - f2).support ⊆ f1.support := by simp only [subset_iff, tsub_eq_zero_iff_le, mem_support_iff, ne.def, coe_tsub, pi.sub_apply, not_imp_not, zero_le, implies_true_iff] {contextual := tt} -lemma subset_support_tsub {f1 f2 : ι →₀ α} : f1.support \ f2.support ⊆ (f1 - f2).support := +lemma subset_support_tsub [decidable_eq ι] {f1 f2 : ι →₀ α} : + f1.support \ f2.support ⊆ (f1 - f2).support := by simp [subset_iff] {contextual := tt} end canonically_ordered_add_monoid @@ -183,6 +185,7 @@ end lemma disjoint_iff {f g : ι →₀ α} : disjoint f g ↔ disjoint f.support g.support := begin + classical, rw [disjoint_iff, disjoint_iff, finsupp.bot_eq_zero, ← finsupp.support_eq_empty, finsupp.support_inf], refl, diff --git a/src/data/finsupp/to_dfinsupp.lean b/src/data/finsupp/to_dfinsupp.lean index 17e7ff470fb9b..9254fc3fa56b0 100644 --- a/src/data/finsupp/to_dfinsupp.lean +++ b/src/data/finsupp/to_dfinsupp.lean @@ -198,7 +198,6 @@ section sigma /-- ### Stronger versions of `finsupp.split` -/ noncomputable theory -open_locale classical variables {η : ι → Type*} {N : Type*} [semiring R] @@ -209,10 +208,12 @@ def sigma_finsupp_equiv_dfinsupp [has_zero N] : ((Σ i, η i) →₀ N) ≃ (Π { to_fun := λ f, ⟨split f, trunc.mk ⟨(split_support f : finset ι).val, λ i, begin rw [← finset.mem_def, mem_split_support_iff_nonzero], - exact (decidable.em _).symm + exact (em _).symm end⟩⟩, inv_fun := λ f, begin + haveI := classical.dec_eq ι, + haveI := λ i, classical.dec_eq (η i →₀ N), refine on_finset (finset.sigma f.support (λ j, (f j).support)) (λ ji, f ji.1 ji.2) (λ g hg, finset.mem_sigma.mpr ⟨_, mem_support_iff.mpr hg⟩), simp only [ne.def, dfinsupp.mem_support_to_fun], @@ -232,7 +233,9 @@ lemma sigma_finsupp_equiv_dfinsupp_symm_apply [has_zero N] (f : Π₀ i, (η i (sigma_finsupp_equiv_dfinsupp.symm f : (Σ i, η i) →₀ N) s = f s.1 s.2 := rfl @[simp] -lemma sigma_finsupp_equiv_dfinsupp_support [has_zero N] (f : (Σ i, η i) →₀ N) : +lemma sigma_finsupp_equiv_dfinsupp_support + [decidable_eq ι] [has_zero N] [Π (i : ι) (x : η i →₀ N), decidable (x ≠ 0)] + (f : (Σ i, η i) →₀ N) : (sigma_finsupp_equiv_dfinsupp f).support = finsupp.split_support f := begin ext, @@ -240,7 +243,8 @@ begin exact (finsupp.mem_split_support_iff_nonzero _ _).symm, end -@[simp] lemma sigma_finsupp_equiv_dfinsupp_single [has_zero N] (a : Σ i, η i) (n : N) : +@[simp] lemma sigma_finsupp_equiv_dfinsupp_single [decidable_eq ι] [has_zero N] + (a : Σ i, η i) (n : N) : sigma_finsupp_equiv_dfinsupp (finsupp.single a n) = @dfinsupp.single _ (λ i, η i →₀ N) _ _ a.1 (finsupp.single a.2 n) := begin @@ -248,6 +252,7 @@ begin ext j b, by_cases h : i = j, { subst h, + classical, simp [split_apply, finsupp.single_apply] }, suffices : finsupp.single (⟨i, a⟩ : Σ i, η i) n ⟨j, b⟩ = 0, { simp [split_apply, dif_neg h, this] }, diff --git a/src/data/mv_polynomial/derivation.lean b/src/data/mv_polynomial/derivation.lean index 161fbbc4fd518..2e88c78ea8837 100644 --- a/src/data/mv_polynomial/derivation.lean +++ b/src/data/mv_polynomial/derivation.lean @@ -110,7 +110,7 @@ def mk_derivation (f : σ → A) : derivation R (mv_polynomial σ R) A := leibniz' := (leibniz_iff_X (mk_derivationₗ R f) (mk_derivationₗ_C _ 1)).2 $ λ s i, begin simp only [mk_derivationₗ_monomial, X, monomial_mul, one_smul, one_mul], - rw [finsupp.sum_add_index]; + rw [finsupp.sum_add_index']; [skip, by simp, by { intros, simp only [nat.cast_add, (monomial _).map_add, add_smul] }], rw [finsupp.sum_single_index, finsupp.sum_single_index]; [skip, by simp, by simp], rw [tsub_self, add_tsub_cancel_right, nat.cast_one, ← C_apply, C_1, one_smul, diff --git a/src/ring_theory/ideal/operations.lean b/src/ring_theory/ideal/operations.lean index 21e13f9e69593..efde096489909 100644 --- a/src/ring_theory/ideal/operations.lean +++ b/src/ring_theory/ideal/operations.lean @@ -257,7 +257,7 @@ begin simp }, { exact ⟨0, λ i, I.zero_mem, finsupp.sum_zero_index⟩ }, { rintros x y ⟨ax, hax, rfl⟩ ⟨ay, hay, rfl⟩, - refine ⟨ax + ay, λ i, I.add_mem (hax i) (hay i), finsupp.sum_add_index _ _⟩; + refine ⟨ax + ay, λ i, I.add_mem (hax i) (hay i), finsupp.sum_add_index' _ _⟩; intros; simp only [zero_smul, add_smul] }, { rintros c x ⟨a, ha, rfl⟩, refine ⟨c • a, λ i, I.mul_mem_left c (ha i), _⟩, From 327c3c0d9232d80e250dc8f65e7835b82b266ea5 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Mon, 23 Jan 2023 12:48:10 +0000 Subject: [PATCH 0331/1291] chore(*): add mathlib4 synchronization comments (#18218) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.big_operators.basic` * `algebra.big_operators.nat_antidiagonal` * `algebra.big_operators.option` * `algebra.big_operators.order` * `algebra.big_operators.pi` * `algebra.big_operators.ring` * `algebra.big_operators.ring_equiv` * `algebra.char_zero.infinite` * `algebra.group.conj_finite` * `algebra.indicator_function` * `algebra.module.big_operators` * `algebra.module.linear_map` * `algebra.order.field.pi` * `algebra.star.big_operators` * `algebra.support` * `combinatorics.additive.energy` * `combinatorics.double_counting` * `combinatorics.quiver.single_obj` * `control.traversable.instances` * `data.W.basic` * `data.finite.basic` * `data.finite.set` * `data.finset.noncomm_prod` * `data.finset.preimage` * `data.finset.sort` * `data.fintype.big_operators` * `data.fintype.card` * `data.fintype.lattice` * `data.fintype.option` * `data.fintype.parity` * `data.fintype.powerset` * `data.fintype.prod` * `data.fintype.small` * `data.fintype.sort` * `data.fintype.sum` * `data.fintype.units` * `data.fintype.vector` * `data.list.fin_range` * `data.list.nodup_equiv_fin` * `data.list.sort` * `data.multiset.sort` * `data.nat.factorial.big_operators` * `data.nat.factors` * `data.nat.fib` * `data.nat.gcd.big_operators` * `data.nat.lattice` * `data.nat.prime_fin` * `data.prod.tprod` * `data.rat.big_operators` * `data.set.finite` * `data.set.pointwise.big_operators` * `data.set.pointwise.list_of_fn` * `data.sym.basic` * `data.vector.basic` * `data.vector.mem` * `group_theory.group_action.big_operators` * `group_theory.submonoid.membership` * `logic.denumerable` * `logic.equiv.list` * `logic.small.list` * `number_theory.frobenius_number` * `order.atoms.finite` * `order.conditionally_complete_lattice.finset` * `order.omega_complete_partial_order` * `ring_theory.coprime.lemmas` * `ring_theory.fintype` * `ring_theory.non_zero_divisors` * `ring_theory.prime` --- src/algebra/big_operators/basic.lean | 3 +++ src/algebra/big_operators/nat_antidiagonal.lean | 3 +++ src/algebra/big_operators/option.lean | 3 +++ src/algebra/big_operators/order.lean | 3 +++ src/algebra/big_operators/pi.lean | 3 +++ src/algebra/big_operators/ring.lean | 3 +++ src/algebra/big_operators/ring_equiv.lean | 3 +++ src/algebra/char_zero/infinite.lean | 5 ++++- src/algebra/group/conj_finite.lean | 3 +++ src/algebra/indicator_function.lean | 3 +++ src/algebra/module/big_operators.lean | 3 +++ src/algebra/module/linear_map.lean | 3 +++ src/algebra/order/field/pi.lean | 3 +++ src/algebra/star/big_operators.lean | 3 +++ src/algebra/support.lean | 3 +++ src/combinatorics/additive/energy.lean | 3 +++ src/combinatorics/double_counting.lean | 3 +++ src/combinatorics/quiver/single_obj.lean | 3 +++ src/control/traversable/instances.lean | 3 +++ src/data/W/basic.lean | 3 +++ src/data/finite/basic.lean | 3 +++ src/data/finite/set.lean | 3 +++ src/data/finset/noncomm_prod.lean | 3 +++ src/data/finset/preimage.lean | 3 +++ src/data/finset/sort.lean | 3 +++ src/data/fintype/big_operators.lean | 3 +++ src/data/fintype/card.lean | 3 +++ src/data/fintype/lattice.lean | 3 +++ src/data/fintype/option.lean | 3 +++ src/data/fintype/parity.lean | 3 +++ src/data/fintype/powerset.lean | 3 +++ src/data/fintype/prod.lean | 3 +++ src/data/fintype/small.lean | 3 +++ src/data/fintype/sort.lean | 3 +++ src/data/fintype/sum.lean | 3 +++ src/data/fintype/units.lean | 3 +++ src/data/fintype/vector.lean | 3 +++ src/data/list/fin_range.lean | 3 +++ src/data/list/nodup_equiv_fin.lean | 3 +++ src/data/list/sort.lean | 3 +++ src/data/multiset/sort.lean | 3 +++ src/data/nat/factorial/big_operators.lean | 3 +++ src/data/nat/factors.lean | 3 +++ src/data/nat/fib.lean | 3 +++ src/data/nat/gcd/big_operators.lean | 3 +++ src/data/nat/lattice.lean | 3 +++ src/data/nat/prime_fin.lean | 3 +++ src/data/prod/tprod.lean | 3 +++ src/data/rat/big_operators.lean | 3 +++ src/data/set/finite.lean | 3 +++ src/data/set/pointwise/big_operators.lean | 3 +++ src/data/set/pointwise/list_of_fn.lean | 3 +++ src/data/sym/basic.lean | 3 +++ src/data/vector/basic.lean | 3 +++ src/data/vector/mem.lean | 3 +++ src/group_theory/group_action/big_operators.lean | 3 +++ src/group_theory/submonoid/membership.lean | 3 +++ src/logic/denumerable.lean | 3 +++ src/logic/equiv/list.lean | 3 +++ src/logic/small/list.lean | 3 +++ src/number_theory/frobenius_number.lean | 3 +++ src/order/atoms/finite.lean | 3 +++ src/order/conditionally_complete_lattice/finset.lean | 3 +++ src/order/omega_complete_partial_order.lean | 3 +++ src/ring_theory/coprime/lemmas.lean | 3 +++ src/ring_theory/fintype.lean | 3 +++ src/ring_theory/non_zero_divisors.lean | 3 +++ src/ring_theory/prime.lean | 3 +++ 68 files changed, 205 insertions(+), 1 deletion(-) diff --git a/src/algebra/big_operators/basic.lean b/src/algebra/big_operators/basic.lean index 9a22cfdda19e7..aa968bc5c9fee 100644 --- a/src/algebra/big_operators/basic.lean +++ b/src/algebra/big_operators/basic.lean @@ -18,6 +18,9 @@ import data.set.pairwise /-! # Big operators +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define products and sums indexed by finite sets (specifically, `finset`). ## Notation diff --git a/src/algebra/big_operators/nat_antidiagonal.lean b/src/algebra/big_operators/nat_antidiagonal.lean index a28d2e4d8eb0a..29f5411651071 100644 --- a/src/algebra/big_operators/nat_antidiagonal.lean +++ b/src/algebra/big_operators/nat_antidiagonal.lean @@ -10,6 +10,9 @@ import algebra.big_operators.basic /-! # Big operators for `nat_antidiagonal` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains theorems relevant to big operators over `finset.nat.antidiagonal`. -/ diff --git a/src/algebra/big_operators/option.lean b/src/algebra/big_operators/option.lean index 2abc0b757aa9c..560824c0e2b79 100644 --- a/src/algebra/big_operators/option.lean +++ b/src/algebra/big_operators/option.lean @@ -10,6 +10,9 @@ import data.finset.option /-! # Lemmas about products and sums over finite sets in `option α` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove formulas for products and sums over `finset.insert_none s` and `finset.erase_none s`. -/ diff --git a/src/algebra/big_operators/order.lean b/src/algebra/big_operators/order.lean index 093e4e390cac3..413f7f89035dc 100644 --- a/src/algebra/big_operators/order.lean +++ b/src/algebra/big_operators/order.lean @@ -12,6 +12,9 @@ import data.fintype.card /-! # Results about big operators with values in an ordered algebraic structure. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Mostly monotonicity results for the `∏` and `∑` operations. -/ diff --git a/src/algebra/big_operators/pi.lean b/src/algebra/big_operators/pi.lean index 3e7a18a02268a..4e6021c97864a 100644 --- a/src/algebra/big_operators/pi.lean +++ b/src/algebra/big_operators/pi.lean @@ -11,6 +11,9 @@ import algebra.ring.pi /-! # Big operators for Pi Types +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains theorems relevant to big operators in binary and arbitrary product of monoids and groups -/ diff --git a/src/algebra/big_operators/ring.lean b/src/algebra/big_operators/ring.lean index 5966de272a8c1..8ed6121b1b0d7 100644 --- a/src/algebra/big_operators/ring.lean +++ b/src/algebra/big_operators/ring.lean @@ -12,6 +12,9 @@ import data.finset.powerset /-! # Results about big operators with values in a (semi)ring +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We prove results about big operators that involve some interaction between multiplicative and additive structures on the values being combined. -/ diff --git a/src/algebra/big_operators/ring_equiv.lean b/src/algebra/big_operators/ring_equiv.lean index 73d0dd0e43425..35864956133cf 100644 --- a/src/algebra/big_operators/ring_equiv.lean +++ b/src/algebra/big_operators/ring_equiv.lean @@ -9,6 +9,9 @@ import algebra.ring.equiv /-! # Results about mapping big operators across ring equivalences + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ namespace ring_equiv diff --git a/src/algebra/char_zero/infinite.lean b/src/algebra/char_zero/infinite.lean index e9cebe06e7997..e204f5bbb4c84 100644 --- a/src/algebra/char_zero/infinite.lean +++ b/src/algebra/char_zero/infinite.lean @@ -6,7 +6,10 @@ Authors: Johan Commelin import algebra.char_zero.defs import data.fintype.card -/-! # A characteristic-zero semiring is infinite -/ +/-! # A characteristic-zero semiring is infinite + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4.-/ open set variables (M : Type*) [add_monoid_with_one M] [char_zero M] diff --git a/src/algebra/group/conj_finite.lean b/src/algebra/group/conj_finite.lean index 44b67a8656686..3329c47cb2c86 100644 --- a/src/algebra/group/conj_finite.lean +++ b/src/algebra/group/conj_finite.lean @@ -10,6 +10,9 @@ import data.fintype.units /-! # Conjugacy of elements of finite groups + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ variables {α : Type*} [monoid α] diff --git a/src/algebra/indicator_function.lean b/src/algebra/indicator_function.lean index cc4c221a06b5f..d412f3c3e71ef 100644 --- a/src/algebra/indicator_function.lean +++ b/src/algebra/indicator_function.lean @@ -8,6 +8,9 @@ import algebra.support /-! # Indicator function +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + - `indicator (s : set α) (f : α → β) (a : α)` is `f a` if `a ∈ s` and is `0` otherwise. - `mul_indicator (s : set α) (f : α → β) (a : α)` is `f a` if `a ∈ s` and is `1` otherwise. diff --git a/src/algebra/module/big_operators.lean b/src/algebra/module/big_operators.lean index 2945a10c102fa..3a64046bf8454 100644 --- a/src/algebra/module/big_operators.lean +++ b/src/algebra/module/big_operators.lean @@ -8,6 +8,9 @@ import group_theory.group_action.big_operators /-! # Finite sums over modules over a ring + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ open_locale big_operators diff --git a/src/algebra/module/linear_map.lean b/src/algebra/module/linear_map.lean index ef23231b5c315..ea7dfe2aa9095 100644 --- a/src/algebra/module/linear_map.lean +++ b/src/algebra/module/linear_map.lean @@ -13,6 +13,9 @@ import algebra.ring.comp_typeclasses /-! # (Semi)linear maps +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define * `linear_map σ M M₂`, `M →ₛₗ[σ] M₂` : a semilinear map between two `module`s. Here, diff --git a/src/algebra/order/field/pi.lean b/src/algebra/order/field/pi.lean index fcbf19e00151f..1b4b66e5a7e25 100644 --- a/src/algebra/order/field/pi.lean +++ b/src/algebra/order/field/pi.lean @@ -9,6 +9,9 @@ import data.fintype.lattice /-! # Lemmas about (finite domain) functions into fields. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We split this from `algebra.order.field.basic` to avoid importing the finiteness hierarchy there. -/ diff --git a/src/algebra/star/big_operators.lean b/src/algebra/star/big_operators.lean index 156f2f8f3ec94..fd310750c5eae 100644 --- a/src/algebra/star/big_operators.lean +++ b/src/algebra/star/big_operators.lean @@ -8,6 +8,9 @@ import algebra.star.basic /-! # Big-operators lemmas about `star` algebraic operations +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + These results are kept separate from `algebra.star.basic` to avoid it needing to import `finset`. -/ diff --git a/src/algebra/support.lean b/src/algebra/support.lean index 79286fe5e7232..f0cdd96b62fca 100644 --- a/src/algebra/support.lean +++ b/src/algebra/support.lean @@ -14,6 +14,9 @@ import group_theory.group_action.pi /-! # Support of a function +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define `function.support f = {x | f x ≠ 0}` and prove its basic properties. We also define `function.mul_support f = {x | f x ≠ 1}`. -/ diff --git a/src/combinatorics/additive/energy.lean b/src/combinatorics/additive/energy.lean index 509c63b16f94f..cb6bbc075dfa7 100644 --- a/src/combinatorics/additive/energy.lean +++ b/src/combinatorics/additive/energy.lean @@ -9,6 +9,9 @@ import data.fintype.prod /-! # Additive energy +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the additive energy of two finsets of a group. This is a central quantity in additive combinatorics. diff --git a/src/combinatorics/double_counting.lean b/src/combinatorics/double_counting.lean index 3c042ba9a7ed4..c500b6f7b8bdf 100644 --- a/src/combinatorics/double_counting.lean +++ b/src/combinatorics/double_counting.lean @@ -8,6 +8,9 @@ import algebra.big_operators.order /-! # Double countings +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file gathers a few double counting arguments. ## Bipartite graphs diff --git a/src/combinatorics/quiver/single_obj.lean b/src/combinatorics/quiver/single_obj.lean index b77762c5b9fb8..8fc315d081ef6 100644 --- a/src/combinatorics/quiver/single_obj.lean +++ b/src/combinatorics/quiver/single_obj.lean @@ -9,6 +9,9 @@ import combinatorics.quiver.symmetric /-! # Single-object quiver +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Single object quiver with a given arrows type. ## Main definitions diff --git a/src/control/traversable/instances.lean b/src/control/traversable/instances.lean index 97fe4a08bf51e..18bd514d9a37c 100644 --- a/src/control/traversable/instances.lean +++ b/src/control/traversable/instances.lean @@ -10,6 +10,9 @@ import data.set.functor /-! # Traversable instances +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file provides instances of `traversable` for types from the core library: `option`, `list` and `sum`. -/ diff --git a/src/data/W/basic.lean b/src/data/W/basic.lean index cd30bc7b23bfb..98e0919dbef24 100644 --- a/src/data/W/basic.lean +++ b/src/data/W/basic.lean @@ -8,6 +8,9 @@ import logic.equiv.list /-! # W types +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given `α : Type` and `β : α → Type`, the W type determined by this data, `W_type β`, is the inductively defined type of trees where the nodes are labeled by elements of `α` and the children of a node labeled `a` are indexed by elements of `β a`. diff --git a/src/data/finite/basic.lean b/src/data/finite/basic.lean index d38068956d901..b92e642b7944d 100644 --- a/src/data/finite/basic.lean +++ b/src/data/finite/basic.lean @@ -12,6 +12,9 @@ import data.fintype.vector /-! # Finite types +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove some theorems about `finite` and provide some instances. This typeclass is a `Prop`-valued counterpart of the typeclass `fintype`. See more details in the file where `finite` is defined. diff --git a/src/data/finite/set.lean b/src/data/finite/set.lean index b46a4d80475da..9436d239c7d16 100644 --- a/src/data/finite/set.lean +++ b/src/data/finite/set.lean @@ -8,6 +8,9 @@ import data.fintype.card /-! # Lemmas about `finite` and `set`s +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove two lemmas about `finite` and `set`s. ## Tags diff --git a/src/data/finset/noncomm_prod.lean b/src/data/finset/noncomm_prod.lean index b6e9fe11fc6cf..cf80a04194924 100644 --- a/src/data/finset/noncomm_prod.lean +++ b/src/data/finset/noncomm_prod.lean @@ -10,6 +10,9 @@ import algebra.big_operators.basic /-! # Products (respectively, sums) over a finset or a multiset. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The regular `finset.prod` and `multiset.prod` require `[comm_monoid α]`. Often, there are collections `s : finset α` where `[monoid α]` and we know, in a dependent fashion, that for all the terms `∀ (x ∈ s) (y ∈ s), commute x y`. diff --git a/src/data/finset/preimage.lean b/src/data/finset/preimage.lean index 18e70d447838e..ba9639d06700d 100644 --- a/src/data/finset/preimage.lean +++ b/src/data/finset/preimage.lean @@ -8,6 +8,9 @@ import algebra.big_operators.basic /-! # Preimage of a `finset` under an injective map. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ open set function diff --git a/src/data/finset/sort.lean b/src/data/finset/sort.lean index c293a1d344bf9..2e7d38585c847 100644 --- a/src/data/finset/sort.lean +++ b/src/data/finset/sort.lean @@ -10,6 +10,9 @@ import data.list.nodup_equiv_fin /-! # Construct a sorted list from a finset. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ namespace finset diff --git a/src/data/fintype/big_operators.lean b/src/data/fintype/big_operators.lean index d2fd5fd156532..c34174eba89f2 100644 --- a/src/data/fintype/big_operators.lean +++ b/src/data/fintype/big_operators.lean @@ -16,6 +16,9 @@ Results about "big operations" over a `fintype`, and consequent results about cardinalities of certain types. ## Implementation note + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. This content had previously been in `data.fintype.basic`, but was moved here to avoid requiring `algebra.big_operators` (and hence many other imports) as a dependency of `fintype`. diff --git a/src/data/fintype/card.lean b/src/data/fintype/card.lean index 9deaf0aa704ef..ab990add5a863 100644 --- a/src/data/fintype/card.lean +++ b/src/data/fintype/card.lean @@ -12,6 +12,9 @@ import tactic.wlog /-! # Cardinalities of finite types +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main declarations * `fintype.card α`: Cardinality of a fintype. Equal to `finset.univ.card`. diff --git a/src/data/fintype/lattice.lean b/src/data/fintype/lattice.lean index 05e0c778330cf..da74d1a887fca 100644 --- a/src/data/fintype/lattice.lean +++ b/src/data/fintype/lattice.lean @@ -8,6 +8,9 @@ import data.finset.lattice /-! # Lemmas relating fintypes and order/lattice structure. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ open function diff --git a/src/data/fintype/option.lean b/src/data/fintype/option.lean index bd3f0ee8d2282..30ba9cda38553 100644 --- a/src/data/fintype/option.lean +++ b/src/data/fintype/option.lean @@ -8,6 +8,9 @@ import data.finset.option /-! # fintype instances for option + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ open function diff --git a/src/data/fintype/parity.lean b/src/data/fintype/parity.lean index a5fe50a158d25..9efe5b21bc594 100644 --- a/src/data/fintype/parity.lean +++ b/src/data/fintype/parity.lean @@ -8,6 +8,9 @@ import algebra.parity /-! # The cardinality of `fin (bit0 n)` is even. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ variables {α : Type*} diff --git a/src/data/fintype/powerset.lean b/src/data/fintype/powerset.lean index 220c0be08d565..00d0d0f68ac52 100644 --- a/src/data/fintype/powerset.lean +++ b/src/data/fintype/powerset.lean @@ -8,6 +8,9 @@ import data.finset.powerset /-! # fintype instance for `set α`, when `α` is a fintype + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ variables {α : Type*} diff --git a/src/data/fintype/prod.lean b/src/data/fintype/prod.lean index e46ad8a5ded57..a307a74dddf4b 100644 --- a/src/data/fintype/prod.lean +++ b/src/data/fintype/prod.lean @@ -9,6 +9,9 @@ import data.finset.prod /-! # fintype instance for the product of two fintypes. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + -/ open function diff --git a/src/data/fintype/small.lean b/src/data/fintype/small.lean index f56a2ef40b2ad..6c24ac71f1cbb 100644 --- a/src/data/fintype/small.lean +++ b/src/data/fintype/small.lean @@ -9,6 +9,9 @@ import logic.small.basic /-! # All finite types are small. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + That is, any `α` with `[fintype α]` is equivalent to a type in any universe. -/ diff --git a/src/data/fintype/sort.lean b/src/data/fintype/sort.lean index 551e297a1002d..5d72701738f36 100644 --- a/src/data/fintype/sort.lean +++ b/src/data/fintype/sort.lean @@ -9,6 +9,9 @@ import data.fintype.basic /-! # Sorting a finite type +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file provides two equivalences for linearly ordered fintypes: * `mono_equiv_of_fin`: Order isomorphism between `α` and `fin (card α)`. * `fin_sum_equiv_of_finset`: Equivalence between `α` and `fin m ⊕ fin n` where `m` and `n` are diff --git a/src/data/fintype/sum.lean b/src/data/fintype/sum.lean index a34e9a08da86e..5359cdd8db246 100644 --- a/src/data/fintype/sum.lean +++ b/src/data/fintype/sum.lean @@ -10,6 +10,9 @@ import logic.embedding.set /-! ## Instances +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We provide the `fintype` instance for the sum of two fintypes. -/ diff --git a/src/data/fintype/units.lean b/src/data/fintype/units.lean index ee6649cc03b96..de94d0ad60d8c 100644 --- a/src/data/fintype/units.lean +++ b/src/data/fintype/units.lean @@ -9,6 +9,9 @@ import data.int.units /-! # fintype instances relating to units + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ variables {α : Type*} diff --git a/src/data/fintype/vector.lean b/src/data/fintype/vector.lean index 9b4b98cf30adb..ab666e8105583 100644 --- a/src/data/fintype/vector.lean +++ b/src/data/fintype/vector.lean @@ -8,6 +8,9 @@ import data.sym.basic /-! # `vector α n` and `sym α n` are fintypes when `α` is. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ variables {α : Type*} diff --git a/src/data/list/fin_range.lean b/src/data/list/fin_range.lean index 0ae1a608ec1a4..cf9423b635b5f 100644 --- a/src/data/list/fin_range.lean +++ b/src/data/list/fin_range.lean @@ -9,6 +9,9 @@ import data.list.perm /-! # Lists of elements of `fin n` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file develops some results on `fin_range n`. -/ diff --git a/src/data/list/nodup_equiv_fin.lean b/src/data/list/nodup_equiv_fin.lean index 052911bb4085a..9ea0ba982d781 100644 --- a/src/data/list/nodup_equiv_fin.lean +++ b/src/data/list/nodup_equiv_fin.lean @@ -10,6 +10,9 @@ import data.list.duplicate /-! # Equivalence between `fin (length l)` and elements of a list +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given a list `l`, * if `l` has no duplicates, then `list.nodup.nth_le_equiv` is the equivalence between diff --git a/src/data/list/sort.lean b/src/data/list/sort.lean index 68fecf73bf0e4..914259bc948b3 100644 --- a/src/data/list/sort.lean +++ b/src/data/list/sort.lean @@ -9,6 +9,9 @@ import data.list.perm /-! # Sorting algorithms on lists +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define `list.sorted r l` to be an alias for `pairwise r l`. This alias is preferred in the case that `r` is a `<` or `≤`-like relation. Then we define two sorting algorithms: `list.insertion_sort` and `list.merge_sort`, and prove their correctness. diff --git a/src/data/multiset/sort.lean b/src/data/multiset/sort.lean index 648d975765165..d5a682e912ad3 100644 --- a/src/data/multiset/sort.lean +++ b/src/data/multiset/sort.lean @@ -8,6 +8,9 @@ import data.multiset.basic /-! # Construct a sorted list from a multiset. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ namespace multiset diff --git a/src/data/nat/factorial/big_operators.lean b/src/data/nat/factorial/big_operators.lean index 0c095d9f8f6b4..e31c5f242280b 100644 --- a/src/data/nat/factorial/big_operators.lean +++ b/src/data/nat/factorial/big_operators.lean @@ -9,6 +9,9 @@ import algebra.big_operators.order /-! # Factorial with big operators +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains some lemmas on factorials in combination with big operators. While in terms of semantics they could be in the `basic.lean` file, importing diff --git a/src/data/nat/factors.lean b/src/data/nat/factors.lean index 09b62ef14243c..6d880604fac04 100644 --- a/src/data/nat/factors.lean +++ b/src/data/nat/factors.lean @@ -11,6 +11,9 @@ import tactic.nth_rewrite /-! # Prime numbers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file deals with the factors of natural numbers. ## Important declarations diff --git a/src/data/nat/fib.lean b/src/data/nat/fib.lean index 84c70569b7456..a56d0fa9c1c7c 100644 --- a/src/data/nat/fib.lean +++ b/src/data/nat/fib.lean @@ -14,6 +14,9 @@ import tactic.wlog /-! # The Fibonacci Sequence +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Summary Definition of the Fibonacci sequence `F₀ = 0, F₁ = 1, Fₙ₊₂ = Fₙ + Fₙ₊₁`. diff --git a/src/data/nat/gcd/big_operators.lean b/src/data/nat/gcd/big_operators.lean index 98c199b463738..3628d43b45b99 100644 --- a/src/data/nat/gcd/big_operators.lean +++ b/src/data/nat/gcd/big_operators.lean @@ -8,6 +8,9 @@ import algebra.big_operators.basic /-! # Lemmas about coprimality with big products. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + These lemmas are kept separate from `data.nat.gcd.basic` in order to minimize imports. -/ diff --git a/src/data/nat/lattice.lean b/src/data/nat/lattice.lean index ff6bbe04e750d..1647489078bf4 100644 --- a/src/data/nat/lattice.lean +++ b/src/data/nat/lattice.lean @@ -8,6 +8,9 @@ import order.conditionally_complete_lattice.finset /-! # Conditionally complete linear order structure on `ℕ` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we * define a `conditionally_complete_linear_order_bot` structure on `ℕ`; diff --git a/src/data/nat/prime_fin.lean b/src/data/nat/prime_fin.lean index 85d8fedbc6a69..9fe7a9f7ef4e1 100644 --- a/src/data/nat/prime_fin.lean +++ b/src/data/nat/prime_fin.lean @@ -10,6 +10,9 @@ import data.set.finite /-! # Prime numbers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains some results about prime numbers which depend on finiteness of sets. -/ diff --git a/src/data/prod/tprod.lean b/src/data/prod/tprod.lean index b89c8772ee701..e7bcd1d83256b 100644 --- a/src/data/prod/tprod.lean +++ b/src/data/prod/tprod.lean @@ -8,6 +8,9 @@ import data.list.nodup /-! # Finite products of types +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the product of types over a list. For `l : list ι` and `α : ι → Type*` we define `list.tprod α l = l.foldr (λ i β, α i × β) punit`. This type should not be used if `Π i, α i` or `Π i ∈ l, α i` can be used instead diff --git a/src/data/rat/big_operators.lean b/src/data/rat/big_operators.lean index d1ca121dbdf3c..faa9503277d6d 100644 --- a/src/data/rat/big_operators.lean +++ b/src/data/rat/big_operators.lean @@ -7,6 +7,9 @@ import data.rat.cast import algebra.big_operators.basic /-! # Casting lemmas for rational numbers involving sums and products + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ open_locale big_operators diff --git a/src/data/set/finite.lean b/src/data/set/finite.lean index 2c444b6b90190..649d62f1ec157 100644 --- a/src/data/set/finite.lean +++ b/src/data/set/finite.lean @@ -10,6 +10,9 @@ import data.finite.basic /-! # Finite sets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines predicates for finite and infinite sets and provides `fintype` instances for many set constructions. It also proves basic facts about finite sets and gives ways to manipulate `set.finite` expressions. diff --git a/src/data/set/pointwise/big_operators.lean b/src/data/set/pointwise/big_operators.lean index 4af1931f000e2..0345699d6e137 100644 --- a/src/data/set/pointwise/big_operators.lean +++ b/src/data/set/pointwise/big_operators.lean @@ -8,6 +8,9 @@ import data.set.pointwise.basic /-! # Results about pointwise operations on sets and big operators. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ namespace set diff --git a/src/data/set/pointwise/list_of_fn.lean b/src/data/set/pointwise/list_of_fn.lean index 3019064cf9ff5..cba70055c1b77 100644 --- a/src/data/set/pointwise/list_of_fn.lean +++ b/src/data/set/pointwise/list_of_fn.lean @@ -10,6 +10,9 @@ import data.list.of_fn /-! # Pointwise operations with lists of sets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves some lemmas about pointwise algebraic operations with lists of sets. -/ diff --git a/src/data/sym/basic.lean b/src/data/sym/basic.lean index f1a70a1fe9a65..7e43f0dc741ad 100644 --- a/src/data/sym/basic.lean +++ b/src/data/sym/basic.lean @@ -12,6 +12,9 @@ import tactic.apply_fun /-! # Symmetric powers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines symmetric powers of a type. The nth symmetric power consists of homogeneous n-tuples modulo permutations by the symmetric group. diff --git a/src/data/vector/basic.lean b/src/data/vector/basic.lean index 8aa74ead6d4e2..f60d184892f4d 100644 --- a/src/data/vector/basic.lean +++ b/src/data/vector/basic.lean @@ -11,6 +11,9 @@ import meta.univs /-! # Additional theorems and definitions about the `vector` type +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file introduces the infix notation `::ᵥ` for `vector.cons`. -/ universes u diff --git a/src/data/vector/mem.lean b/src/data/vector/mem.lean index 7db02e14c7bbe..3a3d4503deb58 100644 --- a/src/data/vector/mem.lean +++ b/src/data/vector/mem.lean @@ -7,6 +7,9 @@ import data.vector.basic /-! # Theorems about membership of elements in vectors +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains theorems for membership in a `v.to_list` for a vector `v`. Having the length available in the type allows some of the lemmas to be simpler and more general than the original version for lists. diff --git a/src/group_theory/group_action/big_operators.lean b/src/group_theory/group_action/big_operators.lean index af25d91dc1244..3c799badb1f36 100644 --- a/src/group_theory/group_action/big_operators.lean +++ b/src/group_theory/group_action/big_operators.lean @@ -11,6 +11,9 @@ import group_theory.group_action.defs /-! # Lemmas about group actions on big operators +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Note that analogous lemmas for `module`s like `finset.sum_smul` appear in other files. -/ diff --git a/src/group_theory/submonoid/membership.lean b/src/group_theory/submonoid/membership.lean index 355e74a4360d6..b51c17fd96fc6 100644 --- a/src/group_theory/submonoid/membership.lean +++ b/src/group_theory/submonoid/membership.lean @@ -12,6 +12,9 @@ import data.finset.noncomm_prod /-! # Submonoids: membership criteria +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove various facts about membership in a submonoid: * `list_prod_mem`, `multiset_prod_mem`, `prod_mem`: if each element of a collection belongs diff --git a/src/logic/denumerable.lean b/src/logic/denumerable.lean index f3510d58ae41f..30c05405be863 100644 --- a/src/logic/denumerable.lean +++ b/src/logic/denumerable.lean @@ -11,6 +11,9 @@ import logic.encodable.basic /-! # Denumerable types +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines denumerable (countably infinite) types as a typeclass extending `encodable`. This is used to provide explicit encode/decode functions from and to `ℕ`, with the information that those functions are inverses of each other. diff --git a/src/logic/equiv/list.lean b/src/logic/equiv/list.lean index 71a3d9e03d6fd..d37847f6caffc 100644 --- a/src/logic/equiv/list.lean +++ b/src/logic/equiv/list.lean @@ -10,6 +10,9 @@ import logic.denumerable /-! # Equivalences involving `list`-like types +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines some additional constructive equivalences using `encodable` and the pairing function on `ℕ`. -/ diff --git a/src/logic/small/list.lean b/src/logic/small/list.lean index eb668f3dcb118..22c0c647cb521 100644 --- a/src/logic/small/list.lean +++ b/src/logic/small/list.lean @@ -9,6 +9,9 @@ import data.vector.basic /-! # Instances for `small (list α)` and `small (vector α)`. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + These must not be in `logic.small.basic` as this is very low in the import hierarchy, and is used by category theory files which do not need everything imported by `data.vector.basic`. -/ diff --git a/src/number_theory/frobenius_number.lean b/src/number_theory/frobenius_number.lean index d740e474a73c8..2a9c1f7a36f1e 100644 --- a/src/number_theory/frobenius_number.lean +++ b/src/number_theory/frobenius_number.lean @@ -12,6 +12,9 @@ import tactic.zify /-! # Frobenius Number in Two Variables +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we first define a predicate for Frobenius numbers, then solve the 2-variable variant of this problem. diff --git a/src/order/atoms/finite.lean b/src/order/atoms/finite.lean index 2a19d6b129053..1f10d0ff97cba 100644 --- a/src/order/atoms/finite.lean +++ b/src/order/atoms/finite.lean @@ -9,6 +9,9 @@ import order.atoms /-! # Atoms, Coatoms, Simple Lattices, and Finiteness +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This module contains some results on atoms and simple lattices in the finite context. ## Main results diff --git a/src/order/conditionally_complete_lattice/finset.lean b/src/order/conditionally_complete_lattice/finset.lean index 182242ff9b5c6..ad588569add27 100644 --- a/src/order/conditionally_complete_lattice/finset.lean +++ b/src/order/conditionally_complete_lattice/finset.lean @@ -9,6 +9,9 @@ import data.set.finite /-! # Conditionally complete lattices and finite sets. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + -/ open set diff --git a/src/order/omega_complete_partial_order.lean b/src/order/omega_complete_partial_order.lean index bf44b4f4c0a47..b75b68bae3ace 100644 --- a/src/order/omega_complete_partial_order.lean +++ b/src/order/omega_complete_partial_order.lean @@ -12,6 +12,9 @@ import tactic.wlog /-! # Omega Complete Partial Orders +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + An omega-complete partial order is a partial order with a supremum operation on increasing sequences indexed by natural numbers (which we call `ωSup`). In this sense, it is strictly weaker than join complete diff --git a/src/ring_theory/coprime/lemmas.lean b/src/ring_theory/coprime/lemmas.lean index c0c3bf5fab27f..6d6c98a25f7ac 100644 --- a/src/ring_theory/coprime/lemmas.lean +++ b/src/ring_theory/coprime/lemmas.lean @@ -11,6 +11,9 @@ import ring_theory.coprime.basic /-! # Additional lemmas about elements of a ring satisfying `is_coprime` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + These lemmas are in a separate file to the definition of `is_coprime` as they require more imports. Notably, this includes lemmas about `finset.prod` as this requires importing big_operators, and diff --git a/src/ring_theory/fintype.lean b/src/ring_theory/fintype.lean index 00047fc067b6b..f5f0a5e97256f 100644 --- a/src/ring_theory/fintype.lean +++ b/src/ring_theory/fintype.lean @@ -7,6 +7,9 @@ import data.fintype.units /-! # Some facts about finite rings + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ open_locale classical diff --git a/src/ring_theory/non_zero_divisors.lean b/src/ring_theory/non_zero_divisors.lean index f7577bfa68dfd..539d03416b96c 100644 --- a/src/ring_theory/non_zero_divisors.lean +++ b/src/ring_theory/non_zero_divisors.lean @@ -10,6 +10,9 @@ import group_theory.submonoid.membership /-! # Non-zero divisors +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define the submonoid `non_zero_divisors` of a `monoid_with_zero`. ## Notations diff --git a/src/ring_theory/prime.lean b/src/ring_theory/prime.lean index 9d37527d15d08..d6b42a534d26d 100644 --- a/src/ring_theory/prime.lean +++ b/src/ring_theory/prime.lean @@ -7,6 +7,9 @@ import algebra.associated import algebra.big_operators.basic /-! # Prime elements in rings + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. This file contains lemmas about prime elements of commutative rings. -/ From 33694ffde850f02f6347ee622b1f766e5a4c908d Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 23 Jan 2023 15:51:25 +0000 Subject: [PATCH 0332/1291] chore(data/finsupp/defs): move lemmas out of the single section which are not about single (#18260) --- src/data/finsupp/defs.lean | 14 +++++++------- 1 file changed, 7 insertions(+), 7 deletions(-) diff --git a/src/data/finsupp/defs.lean b/src/data/finsupp/defs.lean index da86462d707b7..5ddfe84c184a1 100644 --- a/src/data/finsupp/defs.lean +++ b/src/data/finsupp/defs.lean @@ -193,6 +193,13 @@ If `α` has a unique term, the type of finitely supported functions `α →₀ def _root_.equiv.finsupp_unique {ι : Type*} [unique ι] : (ι →₀ M) ≃ M := finsupp.equiv_fun_on_finite.trans (equiv.fun_unique ι M) +@[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 := +⟨λ h, h ▸ rfl, unique_ext⟩ + end basic /-! ### Declarations about `single` -/ @@ -339,13 +346,6 @@ 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 := -⟨λ h, h ▸ rfl, unique_ext⟩ - @[simp] lemma unique_single_eq_iff [unique α] {b' : M} : single a b = single a' b' ↔ b = b' := by rw [unique_ext_iff, unique.eq_default a, unique.eq_default a', single_eq_same, single_eq_same] From 8631e2d5ea77f6c13054d9151d82b83069680cb1 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Mon, 23 Jan 2023 18:46:59 +0000 Subject: [PATCH 0333/1291] refactor(data/nat/parity): reduce imports, add/delete lemmas (#18221) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Sync `mathlib3` with API changes introduced while porting `data.nat.parity` to Mathlib 4 in leanprover-community/mathlib4#1661. * Add `even.two_dvd`, `even.trans_dvd`, `has_dvd.dvd.even`, `odd.of_dvd_nat`, and `odd.ne_two_of_dvd_nat`. * Rename `nat.even_sub_one_of_prime_ne_two` to `nat.prime.even_sub_one`, move to `data.nat.prime`. * Delete `odd.factors_ne_two`. * Replace `is_primitive_root.pow_sub_one_norm_prime_pow_of_one_le` with `is_primitive_root.pow_sub_one_norm_prime_pow_of_ne_zero`, assume `≠ 0` instead of `1 ≤`. --- src/algebra/parity.lean | 13 +++- src/data/bool/count.lean | 1 + src/data/nat/multiplicity.lean | 17 ++--- src/data/nat/parity.lean | 18 ++--- src/data/nat/prime.lean | 3 + .../cyclotomic/discriminant.lean | 67 ++++++++----------- .../cyclotomic/primitive_roots.lean | 8 +-- .../legendre_symbol/jacobi_symbol.lean | 2 +- src/number_theory/primorial.lean | 1 + 9 files changed, 67 insertions(+), 63 deletions(-) diff --git a/src/algebra/parity.lean b/src/algebra/parity.lean index e57fc2b87e6a5..d51bf886d687d 100644 --- a/src/algebra/parity.lean +++ b/src/algebra/parity.lean @@ -167,6 +167,14 @@ by simp [even_iff_exists_two_nsmul] lemma even_iff_two_dvd {a : α} : even a ↔ 2 ∣ a := by simp [even, has_dvd.dvd, two_mul] +alias even_iff_two_dvd ↔ even.two_dvd _ + +theorem even.trans_dvd (hm : even m) (hn : m ∣ n) : even n := +even_iff_two_dvd.2 $ hm.two_dvd.trans hn + +theorem has_dvd.dvd.even (hn : m ∣ n) (hm : even m) : even n := +hm.trans_dvd hn + @[simp] lemma range_two_mul (α : Type*) [semiring α] : set.range (λ x : α, 2 * x) = {a | even a} := by { ext x, simp [eq_comm, two_mul, even] } @@ -214,9 +222,8 @@ lemma odd.add_odd : odd m → odd n → even (m + n) := begin rintro ⟨m, rfl⟩ ⟨n, rfl⟩, refine ⟨n + m + 1, _⟩, - rw [← two_mul, ←add_assoc, add_comm _ (2 * n), ←add_assoc, ←mul_add, add_assoc, mul_add _ (n + m), - mul_one], - refl + rw [two_mul, two_mul], + ac_refl end @[simp] lemma odd_one : odd (1 : α) := diff --git a/src/data/bool/count.lean b/src/data/bool/count.lean index a1d3d07f72ddd..2ac772d0fb6fd 100644 --- a/src/data/bool/count.lean +++ b/src/data/bool/count.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import data.nat.parity +import data.list.chain /-! # List of booleans diff --git a/src/data/nat/multiplicity.lean b/src/data/nat/multiplicity.lean index 4d40db0694f4a..58f3365ed3fa0 100644 --- a/src/data/nat/multiplicity.lean +++ b/src/data/nat/multiplicity.lean @@ -8,6 +8,7 @@ import algebra.geom_sum import data.nat.bitwise import data.nat.log import data.nat.parity +import data.nat.prime import ring_theory.multiplicity /-! @@ -69,21 +70,21 @@ calc namespace prime lemma multiplicity_one {p : ℕ} (hp : p.prime) : multiplicity p 1 = 0 := -multiplicity.one_right (prime_iff.mp hp).not_unit +multiplicity.one_right hp.prime.not_unit lemma multiplicity_mul {p m n : ℕ} (hp : p.prime) : multiplicity p (m * n) = multiplicity p m + multiplicity p n := -multiplicity.mul $ prime_iff.mp hp +multiplicity.mul hp.prime lemma multiplicity_pow {p m n : ℕ} (hp : p.prime) : multiplicity p (m ^ n) = n • (multiplicity p m) := -multiplicity.pow $ prime_iff.mp hp +multiplicity.pow hp.prime lemma multiplicity_self {p : ℕ} (hp : p.prime) : multiplicity p p = 1 := -multiplicity_self (prime_iff.mp hp).not_unit hp.ne_zero +multiplicity_self hp.prime.not_unit hp.ne_zero lemma multiplicity_pow_self {p n : ℕ} (hp : p.prime) : multiplicity p (p ^ n) = n := -multiplicity_pow_self hp.ne_zero (prime_iff.mp hp).not_unit n +multiplicity_pow_self hp.ne_zero hp.prime.not_unit n /-- **Legendre's Theorem** @@ -108,7 +109,7 @@ lemma multiplicity_factorial {p : ℕ} (hp : p.prime) : lemma multiplicity_factorial_mul_succ {n p : ℕ} (hp : p.prime) : multiplicity p (p * (n + 1))! = multiplicity p (p * n)! + multiplicity p (n + 1) + 1 := begin - have hp' := prime_iff.mp hp, + have hp' := hp.prime, have h0 : 2 ≤ p := hp.two_le, have h1 : 1 ≤ p * n + 1 := nat.le_add_left _ _, have h2 : p * n + 1 ≤ p * (n + 1), linarith, @@ -135,7 +136,7 @@ lemma multiplicity_factorial_mul {n p : ℕ} (hp : p.prime) : begin induction n with n ih, { simp }, - { simp only [succ_eq_add_one, multiplicity.mul, hp, prime_iff.mp hp, ih, + { simp only [succ_eq_add_one, multiplicity.mul, hp, hp.prime, ih, multiplicity_factorial_mul_succ, ←add_assoc, nat.cast_one, nat.cast_add, factorial_succ], congr' 1, rw [add_comm, add_assoc] } @@ -227,7 +228,7 @@ end prime lemma multiplicity_two_factorial_lt : ∀ {n : ℕ} (h : n ≠ 0), multiplicity 2 n! < n := begin - have h2 := prime_iff.mp prime_two, + have h2 := prime_two.prime, refine binary_rec _ _, { contradiction }, { intros b n ih h, diff --git a/src/data/nat/parity.lean b/src/data/nat/parity.lean index 88b1ee8dcdd8e..2acd98c48dd03 100644 --- a/src/data/nat/parity.lean +++ b/src/data/nat/parity.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Benjamin Davidson -/ import data.nat.modeq -import data.nat.factors import algebra.parity /-! @@ -191,13 +190,10 @@ begin apply even_mul_succ_self } end -lemma even_sub_one_of_prime_ne_two {p : ℕ} (hp : prime p) (hodd : p ≠ 2) : even (p - 1) := -odd.sub_odd (odd_iff.2 $ hp.eq_two_or_odd.resolve_left hodd) (odd_iff.2 rfl) - lemma two_mul_div_two_of_even : even n → 2 * (n / 2) = n := λ h, nat.mul_div_cancel_left' (even_iff_two_dvd.mp h) -lemma div_two_mul_two_of_even : even n → n / 2 * 2 = n := --nat.div_mul_cancel +lemma div_two_mul_two_of_even : even n → n / 2 * 2 = n := λ h, nat.div_mul_cancel (even_iff_two_dvd.mp h) lemma two_mul_div_two_add_one_of_odd (h : odd n) : 2 * (n / 2) + 1 = n := @@ -296,6 +292,12 @@ lemma odd.mod_even {n a : ℕ} (hn : odd n) (ha : even a) : odd (n % a) := lemma even.mod_even {n a : ℕ} (hn : even n) (ha : even a) : even (n % a) := (even.mod_even_iff ha).mpr hn -/-- `2` is not a prime factor of an odd natural number. -/ -lemma odd.factors_ne_two {n p : ℕ} (hn : odd n) (hp : p ∈ n.factors) : p ≠ 2 := -by { rintro rfl, exact two_dvd_ne_zero.mpr (odd_iff.mp hn) (dvd_of_mem_factors hp) } +theorem odd.of_dvd_nat {m n : ℕ} (hn : odd n) (hm : m ∣ n) : odd m := +odd_iff_not_even.2 $ mt hm.even (odd_iff_not_even.1 hn) + +/-- `2` is not a factor of an odd natural number. -/ +theorem odd.ne_two_of_dvd_nat {m n : ℕ} (hn : odd n) (hm : m ∣ n) : m ≠ 2 := +begin + rintro rfl, + exact absurd (hn.of_dvd_nat hm) dec_trivial +end diff --git a/src/data/nat/prime.lean b/src/data/nat/prime.lean index 57775b325ca8a..ff09a91d35dd7 100644 --- a/src/data/nat/prime.lean +++ b/src/data/nat/prime.lean @@ -447,6 +447,9 @@ by rw [even_iff_two_dvd, prime_dvd_prime_iff_eq prime_two hp, eq_comm] lemma prime.odd_of_ne_two {p : ℕ} (hp : p.prime) (h_two : p ≠ 2) : odd p := hp.eq_two_or_odd'.resolve_left h_two +lemma prime.even_sub_one {p : ℕ} (hp : p.prime) (h2 : p ≠ 2) : even (p - 1) := +let ⟨n, hn⟩ := hp.odd_of_ne_two h2 in ⟨n, by rw [hn, nat.add_sub_cancel, two_mul]⟩ + /-- A prime `p` satisfies `p % 2 = 1` if and only if `p ≠ 2`. -/ lemma prime.mod_two_eq_one_iff_ne_two {p : ℕ} [fact p.prime] : p % 2 = 1 ↔ p ≠ 2 := begin diff --git a/src/number_theory/cyclotomic/discriminant.lean b/src/number_theory/cyclotomic/discriminant.lean index 2a254dc3a5519..03682a366e33f 100644 --- a/src/number_theory/cyclotomic/discriminant.lean +++ b/src/number_theory/cyclotomic/discriminant.lean @@ -67,41 +67,32 @@ lemma discr_prime_pow_ne_two [is_cyclotomic_extension {p ^ (k + 1)} K L] [hp : f (-1) ^ (((p ^ (k + 1) : ℕ).totient) / 2) * p ^ ((p : ℕ) ^ k * ((p - 1) * (k + 1) - 1)) := begin haveI hne := is_cyclotomic_extension.ne_zero' (p ^ (k + 1)) K L, - have hp2 : p = 2 → 1 ≤ k, - { intro hp, - refine one_le_iff_ne_zero.2 (λ h, _), - rw [h, hp, zero_add, pow_one] at hk, - exact hk rfl }, - - rw [discr_power_basis_eq_norm, finrank _ hirr, hζ.power_basis_gen _, - ← hζ.minpoly_eq_cyclotomic_of_irreducible hirr, pnat.pow_coe, totient_prime_pow hp.out - (succ_pos k)], + rw [discr_power_basis_eq_norm, finrank L hirr, hζ.power_basis_gen _, + ← hζ.minpoly_eq_cyclotomic_of_irreducible hirr, pnat.pow_coe, + totient_prime_pow hp.out (succ_pos k), succ_sub_one], + have hp2 : p = 2 → k ≠ 0, + { unfreezingI { rintro rfl rfl }, + exact absurd rfl hk }, congr' 1, - { by_cases hptwo : p = 2, - { obtain ⟨k₁, hk₁⟩ := nat.exists_eq_succ_of_ne_zero (one_le_iff_ne_zero.1 (hp2 hptwo)), - rw [hk₁, succ_sub_one, hptwo, pnat.coe_bit0, pnat.one_coe, succ_sub_succ_eq_sub, tsub_zero, + { unfreezingI { rcases eq_or_ne p 2 with rfl | hp2 }, + { unfreezingI { rcases nat.exists_eq_succ_of_ne_zero (hp2 rfl) with ⟨k, rfl⟩ }, + rw [pnat.coe_bit0, pnat.one_coe, succ_sub_succ_eq_sub, tsub_zero, mul_one, pow_succ, mul_assoc, nat.mul_div_cancel_left _ zero_lt_two, nat.mul_div_cancel_left _ zero_lt_two], - by_cases hk₁zero : k₁ = 0, - { simp [hk₁zero] }, - obtain ⟨k₂, rfl⟩ := nat.exists_eq_succ_of_ne_zero hk₁zero, - rw [pow_succ, mul_assoc, pow_mul (-1 : K), pow_mul (-1 : K), neg_one_sq, one_pow, one_pow] }, - { simp only [succ_sub_succ_eq_sub, tsub_zero], - replace hptwo : ↑p ≠ 2, - { intro h, - rw [← pnat.one_coe, ← pnat.coe_bit0, pnat.coe_inj] at h, - exact hptwo h }, - obtain ⟨a, ha⟩ := even_sub_one_of_prime_ne_two hp.out hptwo, - rw [mul_comm ((p : ℕ) ^ k), mul_assoc, ha], - nth_rewrite 0 [← mul_one a], - nth_rewrite 4 [← mul_one a], - rw [← nat.mul_succ, mul_comm a, mul_assoc, mul_assoc 2, nat.mul_div_cancel_left _ - zero_lt_two, nat.mul_div_cancel_left _ zero_lt_two, ← mul_assoc, mul_comm - (a * (p : ℕ) ^ k), pow_mul, ← ha], - congr' 1, - refine odd.neg_one_pow (nat.even.sub_odd (nat.succ_le_iff.2 (mul_pos (tsub_pos_iff_lt.2 - hp.out.one_lt) (pow_pos hp.out.pos _))) (even.mul_right (nat.even_sub_one_of_prime_ne_two - hp.out hptwo) _) odd_one) } }, + unfreezingI { cases k }, + { simp }, + { rw [pow_succ, (even_two.mul_right _).neg_one_pow, + ((even_two.mul_right _).mul_right _).neg_one_pow] } }, + { replace hp2 : (p : ℕ) ≠ 2, + { rwa [ne.def, ← pnat.one_coe, ← pnat.coe_bit0, pnat.coe_inj] }, + have hpo : odd (p : ℕ) := hp.out.odd_of_ne_two hp2, + obtain ⟨a, ha⟩ := (hp.out.even_sub_one hp2).two_dvd, + rw [ha, mul_left_comm, mul_assoc, nat.mul_div_cancel_left _ two_pos, + nat.mul_div_cancel_left _ two_pos, mul_right_comm, pow_mul, (hpo.pow.mul _).neg_one_pow, + pow_mul, hpo.pow.neg_one_pow], + refine nat.even.sub_odd _ (even_two_mul _) odd_one, + rw [mul_left_comm, ← ha], + exact one_le_mul (one_le_pow _ _ hp.1.pos) (succ_le_iff.2 $ tsub_pos_of_lt hp.1.one_lt) } }, { have H := congr_arg derivative (cyclotomic_prime_pow_mul_X_pow_sub_one K p k), rw [derivative_mul, derivative_sub, derivative_one, sub_zero, derivative_X_pow, C_eq_nat_cast, derivative_sub, derivative_one, sub_zero, derivative_X_pow, C_eq_nat_cast, ← pnat.pow_coe, @@ -112,22 +103,20 @@ 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 rfl.le (hp2 hp) }, - { exact hζ.pow_sub_one_norm_prime_ne_two hirr rfl.le hp } }, + { exact hζ.pow_sub_one_norm_prime_pow_of_ne_zero hirr le_rfl (hp2 hp) }, + { exact hζ.pow_sub_one_norm_prime_ne_two hirr le_rfl 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), + algebra.norm_algebra_map, finrank L 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, hζ.norm_eq_one hk hirr, one_pow, mul_one, cast_pow, ← coe_coe, ← pow_mul, ← mul_assoc, mul_comm (k + 1), mul_assoc] at H, - { have := mul_pos (succ_pos k) (tsub_pos_iff_lt.2 hp.out.one_lt), + { have := mul_pos (succ_pos k) (tsub_pos_of_lt hp.out.one_lt), rw [← succ_pred_eq_of_pos this, mul_succ, pow_add _ _ ((p : ℕ) ^ k)] at H, replace H := (mul_left_inj' (λ h, _)).1 H, { simpa only [← pnat.pow_coe, H, mul_comm _ (k + 1)] }, { replace h := pow_eq_zero h, rw [coe_coe] at h, - simpa using hne.1 } }, - all_goals { apply_instance } }, - all_goals { apply_instance } + simpa using hne.1 } } } end /-- If `p` is a prime and `is_cyclotomic_extension {p ^ (k + 1)} K L`, then the discriminant of diff --git a/src/number_theory/cyclotomic/primitive_roots.lean b/src/number_theory/cyclotomic/primitive_roots.lean index a2865d33a9a1e..802553dac2667 100644 --- a/src/number_theory/cyclotomic/primitive_roots.lean +++ b/src/number_theory/cyclotomic/primitive_roots.lean @@ -442,11 +442,11 @@ begin end /-- 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 {k s : ℕ} (hζ : is_primitive_root ζ ↑(p ^ (k + 1))) +then the norm of `ζ ^ (p ^ s) - 1` is `p ^ (p ^ s)` if `k ≠ 0` and `s ≤ k`. -/ +lemma pow_sub_one_norm_prime_pow_of_ne_zero {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)) (hs : s ≤ k) - (hk : 1 ≤ k) : norm K (ζ ^ ((p : ℕ) ^ s) - 1) = p ^ ((p : ℕ) ^ s) := + (hk : k ≠ 0) : norm K (ζ ^ ((p : ℕ) ^ s) - 1) = p ^ ((p : ℕ) ^ s) := begin by_cases htwo : p ^ (k - s + 1) = 2, { have hp : p = 2, @@ -463,7 +463,7 @@ begin simp only [hs, hp, pnat.coe_bit0, one_coe, coe_coe, cast_bit0, cast_one, pow_coe] at ⊢ hζ hirr hcycl, haveI := hcycl, - obtain ⟨k₁, hk₁⟩ := nat.exists_eq_succ_of_ne_zero (one_le_iff_ne_zero.1 hk), + obtain ⟨k₁, hk₁⟩ := nat.exists_eq_succ_of_ne_zero hk, 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] }, diff --git a/src/number_theory/legendre_symbol/jacobi_symbol.lean b/src/number_theory/legendre_symbol/jacobi_symbol.lean index d4ada59e8e139..9bbec60ab8148 100644 --- a/src/number_theory/legendre_symbol/jacobi_symbol.lean +++ b/src/number_theory/legendre_symbol/jacobi_symbol.lean @@ -244,7 +244,7 @@ begin conv_rhs { rw [← prod_factors hb.pos.ne', cast_list_prod, χ.map_list_prod] }, rw [jacobi_sym, list.map_map, ← list.pmap_eq_map nat.prime _ _ (λ _, prime_of_mem_factors)], congr' 1, apply list.pmap_congr, - exact λ p h pp _, hp p pp (hb.factors_ne_two h), + exact λ p h pp _, hp p pp (hb.ne_two_of_dvd_nat $ dvd_of_mem_factors h) end /-- If `b` is odd, then `J(-1 | b)` is given by `χ₄ b`. -/ diff --git a/src/number_theory/primorial.lean b/src/number_theory/primorial.lean index bcb4485006754..9e8a6174d748d 100644 --- a/src/number_theory/primorial.lean +++ b/src/number_theory/primorial.lean @@ -6,6 +6,7 @@ Authors: Patrick Stevens import algebra.big_operators.associated import data.nat.choose.sum import data.nat.parity +import data.nat.prime import tactic.ring_exp /-! From 808ea4ebfabeb599f21ec4ae87d6dc969597887f Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Mon, 23 Jan 2023 22:33:24 +0000 Subject: [PATCH 0334/1291] feat(algebraic_geometry/elliptic_curve/weierstrass): define a basis for the coordinate ring (#18101) Also refactor `coordinate_ring` definitions and lemmas into the `coordinate_ring` namespace, and compute the degree of the norm of an element of the `coordinate_ring` as an `R[X]`-algebra in terms of this basis. --- .../elliptic_curve/weierstrass.lean | 176 ++++++++++++++++-- src/data/polynomial/degree/definitions.lean | 5 +- 2 files changed, 164 insertions(+), 17 deletions(-) diff --git a/src/algebraic_geometry/elliptic_curve/weierstrass.lean b/src/algebraic_geometry/elliptic_curve/weierstrass.lean index 15703ad46c5da..791d84b50b381 100644 --- a/src/algebraic_geometry/elliptic_curve/weierstrass.lean +++ b/src/algebraic_geometry/elliptic_curve/weierstrass.lean @@ -5,7 +5,7 @@ Authors: Kevin Buzzard, David Kurniadi Angdinata -/ import algebra.cubic_discriminant -import ring_theory.class_group +import ring_theory.norm import tactic.linear_combination /-! @@ -39,6 +39,7 @@ splitting field of `R` are precisely the $X$-coordinates of the non-zero 2-torsi * `weierstrass_curve.nonsingular`: the nonsingular condition at a point on a Weierstrass curve. * `weierstrass_curve.coordinate_ring`: the coordinate ring of a Weierstrass curve. * `weierstrass_curve.function_field`: the function field of a Weierstrass curve. + * `weierstrass_curve.basis`: the power basis of the coordinate ring as an `R[X]`-algebra. * `elliptic_curve`: an elliptic curve over a commutative ring. * `elliptic_curve.j`: the j-invariant of an elliptic curve. @@ -50,6 +51,8 @@ splitting field of `R` are precisely the $X$-coordinates of the non-zero 2-torsi if its discriminant is non-zero. * `weierstrass_curve.coordinate_ring.is_domain`: the coordinate ring of a Weierstrass curve is an integral domain. + * `weierstrass_curve.degree_norm_smul_basis`: the degree of the norm of an element in the + coordinate ring as an `R[X]`-algebra in terms of the power basis. * `elliptic_curve.nonsingular`: an elliptic curve is nonsingular at every point. * `elliptic_curve.variable_change_j`: the j-invariant of an elliptic curve is invariant under an admissible linear change of variables. @@ -80,6 +83,8 @@ private meta def map_simp : tactic unit := private meta def eval_simp : tactic unit := `[simp only [eval_C, eval_X, eval_add, eval_sub, eval_mul, eval_pow]] +private meta def C_simp : tactic unit := `[simp only [C_0, C_1, C_neg, C_add, C_sub, C_mul, C_pow]] + universes u v variable {R : Type u} @@ -251,21 +256,21 @@ X ^ 2 + C (C W.a₁ * X + C W.a₃) * X - C (X ^ 3 + C W.a₂ * X ^ 2 + C W.a₄ lemma polynomial_eq : W.polynomial = cubic.to_poly ⟨0, 1, cubic.to_poly ⟨0, 0, W.a₁, W.a₃⟩, cubic.to_poly ⟨-1, -W.a₂, -W.a₄, -W.a₆⟩⟩ := -by { simp only [weierstrass_curve.polynomial, cubic.to_poly, C_0, C_1, C_neg, C_add, C_mul], ring1 } +by { simp only [weierstrass_curve.polynomial, cubic.to_poly], C_simp, ring1 } lemma polynomial_ne_zero [nontrivial R] : W.polynomial ≠ 0 := by { rw [polynomial_eq], exact cubic.ne_zero_of_b_ne_zero one_ne_zero } -lemma degree_polynomial [nontrivial R] : W.polynomial.degree = 2 := +@[simp] lemma degree_polynomial [nontrivial R] : W.polynomial.degree = 2 := by { rw [polynomial_eq], exact cubic.degree_of_b_ne_zero' one_ne_zero } -lemma nat_degree_polynomial [nontrivial R] : W.polynomial.nat_degree = 2 := +@[simp] lemma nat_degree_polynomial [nontrivial R] : W.polynomial.nat_degree = 2 := by { rw [polynomial_eq], exact cubic.nat_degree_of_b_ne_zero' one_ne_zero } lemma monic_polynomial : W.polynomial.monic := by { nontriviality R, simpa only [polynomial_eq] using cubic.monic_of_b_eq_one' } -lemma irreducible_polynomial [nontrivial R] [no_zero_divisors R] : irreducible W.polynomial := +lemma irreducible_polynomial [is_domain R] : irreducible W.polynomial := begin by_contra h, rcases (W.monic_polynomial.not_irreducible_iff_exists_add_mul_eq_coeff W.nat_degree_polynomial).mp @@ -367,7 +372,7 @@ lemma nonsingular_of_Δ_ne_zero {x y : R} (h : W.equation x y) (hΔ : W.Δ ≠ 0 nonsingular_zero_of_Δ_ne_zero _ ((W.equation_iff_variable_change x y).mp h) $ by rwa [variable_change_Δ, inv_one, units.coe_one, one_pow, one_mul] -/-! ### The coordinate ring -/ +/-! ### Ideals in the coordinate ring -/ /-- The coordinate ring $R[W] := R[X, Y] / \langle W(X, Y) \rangle$ of `W`. @@ -382,39 +387,178 @@ See Zulip thread: https://leanprover.zulipchat.com/#narrow/stream/116395-maths/topic/.E2.9C.94.20class_group.2Emk -/ @[derive [inhabited, comm_ring]] def coordinate_ring : Type u := adjoin_root W.polynomial +/-- The function field $R(W) := \mathrm{Frac}(R[W])$ of `W`. -/ +@[reducible] def function_field : Type u := fraction_ring W.coordinate_ring + +namespace coordinate_ring + instance [is_domain R] [normalized_gcd_monoid R] : is_domain W.coordinate_ring := (ideal.quotient.is_domain_iff_prime _).mpr $ by simpa only [ideal.span_singleton_prime W.polynomial_ne_zero, ← gcd_monoid.irreducible_iff_prime] using W.irreducible_polynomial -instance coordinate_ring.is_domain_of_field {F : Type u} [field F] (W : weierstrass_curve F) : +instance is_domain_of_field {F : Type u} [field F] (W : weierstrass_curve F) : is_domain W.coordinate_ring := by { classical, apply_instance } -/-- The function field $R(W) := \mathrm{Frac}(R[W])$ of `W`. -/ -@[reducible] def function_field : Type u := fraction_ring W.coordinate_ring - variables (x : R) (y : R[X]) /-- The class of the element $X - x$ in $R[W]$ for some $x \in R$. -/ @[simp] noncomputable def X_class : W.coordinate_ring := adjoin_root.mk W.polynomial $ C $ X - C x -lemma X_class_ne_zero [nontrivial R] : W.X_class x ≠ 0 := -adjoin_root.mk_ne_zero_of_nat_degree_lt W.monic_polynomial (C_ne_zero.2 $ X_sub_C_ne_zero x) $ +lemma X_class_ne_zero [nontrivial R] : X_class W x ≠ 0 := +adjoin_root.mk_ne_zero_of_nat_degree_lt W.monic_polynomial (C_ne_zero.mpr $ X_sub_C_ne_zero x) $ by { rw [nat_degree_polynomial, nat_degree_C], norm_num1 } /-- The class of the element $Y - y(X)$ in $R[W]$ for some $y(X) \in R[X]$. -/ @[simp] noncomputable def Y_class : W.coordinate_ring := adjoin_root.mk W.polynomial $ X - C y -lemma Y_class_ne_zero [nontrivial R] : W.Y_class y ≠ 0 := -adjoin_root.mk_ne_zero_of_nat_degree_lt W.monic_polynomial (X_sub_C_ne_zero _) $ +lemma Y_class_ne_zero [nontrivial R] : Y_class W y ≠ 0 := +adjoin_root.mk_ne_zero_of_nat_degree_lt W.monic_polynomial (X_sub_C_ne_zero y) $ by { rw [nat_degree_polynomial, nat_degree_X_sub_C], norm_num1 } /-- The ideal $\langle X - x \rangle$ of $R[W]$ for some $x \in R$. -/ -@[simp] noncomputable def X_ideal : ideal W.coordinate_ring := ideal.span {W.X_class x} +@[simp] noncomputable def X_ideal : ideal W.coordinate_ring := ideal.span {X_class W x} /-- The ideal $\langle Y - y(X) \rangle$ of $R[W]$ for some $y(X) \in R[X]$. -/ -@[simp] noncomputable def Y_ideal : ideal W.coordinate_ring := ideal.span {W.Y_class y} +@[simp] noncomputable def Y_ideal : ideal W.coordinate_ring := ideal.span {Y_class W y} + +/-! ### The coordinate ring as an `R[X]`-algebra -/ + +noncomputable instance : algebra R[X] W.coordinate_ring := ideal.quotient.algebra R[X] + +noncomputable instance algebra' : algebra R W.coordinate_ring := ideal.quotient.algebra R + +instance : is_scalar_tower R R[X] W.coordinate_ring := ideal.quotient.is_scalar_tower R R[X] _ + +instance [subsingleton R] : subsingleton W.coordinate_ring := module.subsingleton R[X] _ + +/-- The basis $\{1, Y\}$ for the coordinate ring $R[W]$ over the polynomial ring $R[X]$. + +Given a Weierstrass curve `W`, write `W^.coordinate_ring.basis` for this basis. -/ +protected noncomputable def basis : basis (fin 2) R[X] W.coordinate_ring := +(subsingleton_or_nontrivial R).by_cases (λ _, by exactI default) $ λ _, by exactI + (basis.reindex (adjoin_root.power_basis' W.monic_polynomial).basis $ + fin_congr $ W.nat_degree_polynomial) + +lemma basis_apply (n : fin 2) : + W^.coordinate_ring.basis n = (adjoin_root.power_basis' W.monic_polynomial).gen ^ (n : ℕ) := +by { nontriviality R, simpa only [coordinate_ring.basis, or.by_cases, dif_neg (not_subsingleton R), + basis.reindex_apply, power_basis.basis_eq_pow] } + +lemma basis_zero : W^.coordinate_ring.basis 0 = 1 := by simpa only [basis_apply] using pow_zero _ + +lemma basis_one : W^.coordinate_ring.basis 1 = adjoin_root.mk W.polynomial X := +by simpa only [basis_apply] using pow_one _ + +@[simp] lemma coe_basis : + (W^.coordinate_ring.basis : fin 2 → W.coordinate_ring) = ![1, adjoin_root.mk W.polynomial X] := +by { ext n, fin_cases n, exacts [basis_zero W, basis_one W] } + +variable {W} + +lemma smul (x : R[X]) (y : W.coordinate_ring) : x • y = adjoin_root.mk W.polynomial (C x) * y := +(algebra_map_smul W.coordinate_ring x y).symm + +lemma smul_basis_eq_zero {p q : R[X]} + (hpq : p • 1 + q • adjoin_root.mk W.polynomial X = 0) : p = 0 ∧ q = 0 := +begin + have h := fintype.linear_independent_iff.mp (coordinate_ring.basis W).linear_independent ![p, q], + erw [fin.sum_univ_succ, basis_zero, fin.sum_univ_one, basis_one] at h, + exact ⟨h hpq 0, h hpq 1⟩ +end + +lemma exists_smul_basis_eq (x : W.coordinate_ring) : + ∃ p q : R[X], p • 1 + q • adjoin_root.mk W.polynomial X = x := +begin + have h := (coordinate_ring.basis W).sum_equiv_fun x, + erw [fin.sum_univ_succ, fin.sum_univ_one, basis_zero, basis_one] at h, + exact ⟨_, _, h⟩ +end + +variable (W) + +lemma smul_basis_mul_C (p q : R[X]) : + (p • 1 + q • adjoin_root.mk W.polynomial X) * adjoin_root.mk W.polynomial (C y) + = ((p * y) • 1 + (q * y) • adjoin_root.mk W.polynomial X) := +by { simp only [smul, map_mul], ring1 } + +lemma smul_basis_mul_Y (p q : R[X]) : + (p • 1 + q • adjoin_root.mk W.polynomial X) * adjoin_root.mk W.polynomial X + = (q * (X ^ 3 + C W.a₂ * X ^ 2 + C W.a₄ * X + C W.a₆)) • 1 + + (p - q * (C W.a₁ * X + C W.a₃)) • adjoin_root.mk W.polynomial X := +begin + have Y_sq : adjoin_root.mk W.polynomial X ^ 2 = adjoin_root.mk W.polynomial + (C (X ^ 3 + C W.a₂ * X ^ 2 + C W.a₄ * X + C W.a₆) - C (C W.a₁ * X + C W.a₃) * X) := + adjoin_root.mk_eq_mk.mpr ⟨1, by { simp only [weierstrass_curve.polynomial], ring1 }⟩, + simp only [smul, add_mul, mul_assoc, ← sq, Y_sq, map_sub, map_mul], + ring1 +end + +/-! ### Norms on the coordinate ring -/ + +lemma norm_smul_basis (p q : R[X]) : + algebra.norm R[X] (p • 1 + q • adjoin_root.mk W.polynomial X) + = p ^ 2 - p * q * (C W.a₁ * X + C W.a₃) + - q ^ 2 * (X ^ 3 + C W.a₂ * X ^ 2 + C W.a₄ * X + C W.a₆) := +begin + simp_rw [algebra.norm_eq_matrix_det W^.coordinate_ring.basis, matrix.det_fin_two, + algebra.left_mul_matrix_eq_repr_mul, basis_zero, mul_one, basis_one, smul_basis_mul_Y, + map_add, finsupp.add_apply, map_smul, finsupp.smul_apply, ← basis_zero, ← basis_one, + basis.repr_self_apply, if_pos, if_neg one_ne_zero, if_neg zero_ne_one, smul_eq_mul], + ring1 +end + +lemma coe_norm_smul_basis (p q : R[X]) : + ↑(algebra.norm R[X] $ p • 1 + q • adjoin_root.mk W.polynomial X) + = adjoin_root.mk W.polynomial + ((C p + C q * X) * (C p + C q * (-X - C (C W.a₁ * X + C W.a₃)))) := +adjoin_root.mk_eq_mk.mpr + ⟨C q ^ 2, by { rw [norm_smul_basis, weierstrass_curve.polynomial], C_simp, ring1 }⟩ + +lemma degree_norm_smul_basis [is_domain R] (p q : R[X]) : + (algebra.norm R[X] $ p • 1 + q • adjoin_root.mk W.polynomial X).degree + = max (2 • p.degree) (2 • q.degree + 3) := +begin + have hdp : (p ^ 2).degree = 2 • p.degree := degree_pow p 2, + have hdpq : (p * q * (C W.a₁ * X + C W.a₃)).degree ≤ p.degree + q.degree + 1, + { simpa only [degree_mul] using add_le_add_left degree_linear_le (p.degree + q.degree) }, + have hdq : (q ^ 2 * (X ^ 3 + C W.a₂ * X ^ 2 + C W.a₄ * X + C W.a₆)).degree = 2 • q.degree + 3, + { rw [degree_mul, degree_pow, ← one_mul $ X ^ 3, ← C_1, degree_cubic $ one_ne_zero' R] }, + rw [norm_smul_basis], + by_cases hp : p = 0, { simpa only [hp, hdq, neg_zero, zero_sub, zero_mul, zero_pow zero_lt_two, + degree_neg] using (max_bot_left _).symm }, + by_cases hq : q = 0, { simpa only [hq, hdp, sub_zero, zero_mul, mul_zero, zero_pow zero_lt_two] + using (max_bot_right _).symm }, + rw [← not_iff_not_of_iff degree_eq_bot] at hp hq, + cases p.degree with dp, { exact (hp rfl).elim }, + cases q.degree with dq, { exact (hq rfl).elim }, + cases le_or_lt dp (dq + 1) with hpq hpq, + { convert (degree_sub_eq_right_of_degree_lt $ (degree_sub_le _ _).trans_lt $ + max_lt_iff.mpr ⟨hdp.trans_lt _, hdpq.trans_lt _⟩).trans (max_eq_right_of_lt _).symm; rw [hdq]; + exact with_bot.coe_lt_coe.mpr (by linarith only [hpq]) }, + { rw [sub_sub], + convert (degree_sub_eq_left_of_degree_lt $ (degree_add_le _ _).trans_lt $ + max_lt_iff.mpr ⟨hdpq.trans_lt _, hdq.trans_lt _⟩).trans (max_eq_left_of_lt _).symm; rw [hdp]; + exact with_bot.coe_lt_coe.mpr (by linarith only [hpq]) } +end + +variable {W} + +lemma degree_norm_ne_one [is_domain R] (x : W.coordinate_ring) : (algebra.norm R[X] x).degree ≠ 1 := +begin + rcases exists_smul_basis_eq x with ⟨p, q, rfl⟩, + rw [degree_norm_smul_basis], + rcases p.degree with (_ | _ | _ | _); cases q.degree, + any_goals { rintro (_ | _) }, + exact (lt_max_of_lt_right dec_trivial).ne' +end + +lemma nat_degree_norm_ne_one [is_domain R] (x : W.coordinate_ring) : + (algebra.norm R[X] x).nat_degree ≠ 1 := +mt (degree_eq_iff_nat_degree_eq_of_pos zero_lt_one).mpr $ degree_norm_ne_one x + +end coordinate_ring end polynomial diff --git a/src/data/polynomial/degree/definitions.lean b/src/data/polynomial/degree/definitions.lean index bac87138058bb..3aee2dc61ef71 100644 --- a/src/data/polynomial/degree/definitions.lean +++ b/src/data/polynomial/degree/definitions.lean @@ -1038,7 +1038,10 @@ section ring variables [ring R] {p q : R[X]} lemma degree_sub_le (p q : R[X]) : degree (p - q) ≤ max (degree p) (degree q) := -by simpa only [sub_eq_add_neg, degree_neg q] using degree_add_le p (-q) +by simpa only [degree_neg q] using degree_add_le p (-q) + +lemma nat_degree_sub_le (p q : R[X]) : nat_degree (p - q) ≤ max (nat_degree p) (nat_degree q) := +by simpa only [← nat_degree_neg q] using nat_degree_add_le p (-q) lemma degree_sub_lt (hd : degree p = degree q) (hp0 : p ≠ 0) (hlc : leading_coeff p = leading_coeff q) : From 68626536f524fb5f7f6c16154d0993e5780cff8e Mon Sep 17 00:00:00 2001 From: Kyle Miller Date: Tue, 24 Jan 2023 00:38:49 +0000 Subject: [PATCH 0335/1291] feat(linear_algebra/dual): more about dual annihilator, some refactoring, and documentation (#17521) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit * Extend documentation with descriptions of main definitions and results. * Move `linear_map.dual_map` earlier in the file and redefine it to use `module.transpose`. * `fintype` -> `set.finite` for `module.dual_bases`. * Rename `submodule.dual_annihilator_comap` to `submodule.dual_coannihilator`. * Add galois correspondence and coinsertion for `submodule.dual_annihilator`; replace proofs using standard GC lemmas. * Add `submodule.dual_pairing` and `submodule.dual_copairing` for bilinear forms that are nondegenerate for subspaces. * Add missing isomorphisms associated to these. * Generalize `f.dual_map.range = f.ker.dual_annihilator` to vector spaces of arbitrary dimension. * Add lemmas for facts such as `(W ⊓ W').dual_annihilator = W.dual_annihilator ⊔ W'.dual_annihilator`. Thanks to Junyan Xu for the generalization to `dual_quot_equiv_dual_annihilator`. --- src/linear_algebra/basic.lean | 5 + src/linear_algebra/bilinear_form.lean | 12 +- src/linear_algebra/dual.lean | 809 ++++++++++++++++++-------- 3 files changed, 575 insertions(+), 251 deletions(-) diff --git a/src/linear_algebra/basic.lean b/src/linear_algebra/basic.lean index ba3b23ccbe75f..00493cd4349f5 100644 --- a/src/linear_algebra/basic.lean +++ b/src/linear_algebra/basic.lean @@ -1472,6 +1472,11 @@ end linear_map f.range_restrict.range = ⊤ := by simp [f.range_cod_restrict _] +@[simp] lemma linear_map.ker_range_restrict [semiring R] [add_comm_monoid M] + [add_comm_monoid M₂] [module R M] [module R M₂] (f : M →ₗ[R] M₂) : + f.range_restrict.ker = f.ker := +linear_map.ker_cod_restrict _ _ _ + /-! ### Linear equivalences -/ namespace linear_equiv diff --git a/src/linear_algebra/bilinear_form.lean b/src/linear_algebra/bilinear_form.lean index c297d997a7f2d..eb4d411cfb3bf 100644 --- a/src/linear_algebra/bilinear_form.lean +++ b/src/linear_algebra/bilinear_form.lean @@ -1090,15 +1090,15 @@ begin exact hx.2 _ submodule.mem_top } end -lemma to_lin_restrict_range_dual_annihilator_comap_eq_orthogonal +lemma to_lin_restrict_range_dual_coannihilator_eq_orthogonal (B : bilin_form K V) (W : subspace K V) : - (B.to_lin.dom_restrict W).range.dual_annihilator_comap = B.orthogonal W := + (B.to_lin.dom_restrict W).range.dual_coannihilator = B.orthogonal W := begin ext x, split; rw [mem_orthogonal_iff]; intro hx, { intros y hy, - rw submodule.mem_dual_annihilator_comap at hx, + rw submodule.mem_dual_coannihilator at hx, refine hx (B.to_lin.dom_restrict W ⟨y, hy⟩) ⟨⟨y, hy⟩, rfl⟩ }, - { rw submodule.mem_dual_annihilator_comap, + { rw submodule.mem_dual_coannihilator, rintro _ ⟨⟨w, hw⟩, rfl⟩, exact hx w hw } end @@ -1113,9 +1113,9 @@ lemma finrank_add_finrank_orthogonal finrank K V + finrank K (W ⊓ B.orthogonal ⊤ : subspace K V) := begin rw [← to_lin_restrict_ker_eq_inf_orthogonal _ _ b₁, - ← to_lin_restrict_range_dual_annihilator_comap_eq_orthogonal _ _, + ← to_lin_restrict_range_dual_coannihilator_eq_orthogonal _ _, finrank_map_subtype_eq], - conv_rhs { rw [← @subspace.finrank_add_finrank_dual_annihilator_comap_eq K V _ _ _ _ + conv_rhs { rw [← @subspace.finrank_add_finrank_dual_coannihilator_eq K V _ _ _ _ (B.to_lin.dom_restrict W).range, add_comm, ← add_assoc, add_comm (finrank K ↥((B.to_lin.dom_restrict W).ker)), linear_map.finrank_range_add_finrank_ker] }, diff --git a/src/linear_algebra/dual.lean b/src/linear_algebra/dual.lean index 632e5effca548..774b9fb9ad784 100644 --- a/src/linear_algebra/dual.lean +++ b/src/linear_algebra/dual.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2019 Johan Commelin. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Johan Commelin, Fabian Glöckle +Authors: Johan Commelin, Fabian Glöckle, Kyle Miller -/ import linear_algebra.finite_dimensional import linear_algebra.projection @@ -12,27 +12,74 @@ import linear_algebra.free_module.finite.basic /-! # Dual vector spaces -The dual space of an R-module M is the R-module of linear maps `M → R`. +The dual space of an $R$-module $M$ is the $R$-module of $R$-linear maps $M \to R$. ## Main definitions -* `dual R M` defines the dual space of M over R. -* Given a basis for an `R`-module `M`, `basis.to_dual` produces a map from `M` to `dual R M`. -* Given families of vectors `e` and `ε`, `module.dual_bases e ε` states that these families have the - characteristic properties of a basis and a dual. -* `dual_annihilator W` is the submodule of `dual R M` where every element annihilates `W`. +* Duals and transposes: + * `module.dual R M` defines the dual space of the `R`-module `M`, as `M →ₗ[R] R`. + * `module.dual_pairing R M` is the canonical pairing between `dual R M` and `M`. + * `module.dual.eval R M : M →ₗ[R] dual R (dual R)` is the canonical map to the double dual. + * `module.dual.transpose` is the linear map from `M →ₗ[R] M'` to `dual R M' →ₗ[R] dual R M`. + * `linear_map.dual_map` is `module.dual.transpose` of a given linear map, for dot notation. + * `linear_equiv.dual_map` is for the dual of an equivalence. +* Bases: + * `basis.to_dual` produces the map `M →ₗ[R] dual R M` associated to a basis for an `R`-module `M`. + * `basis.to_dual_equiv` is the equivalence `M ≃ₗ[R] dual R M` associated to a finite basis. + * `basis.dual_basis` is a basis for `dual R M` given a finite basis for `M`. + * `module.dual_bases e ε` is the proposition that the families `e` of vectors and `ε` of dual + vectors have the characteristic properties of a basis and a dual. +* Submodules: + * `submodule.dual_restrict W` is the transpose `dual R M →ₗ[R] dual R W` of the inclusion map. + * `submodule.dual_annihilator W` is the kernel of `W.dual_restrict`. That is, it is the submodule + of `dual R M` whose elements all annihilate `W`. + * `submodule.dual_restrict_comap W'` is the dual annihilator of `W' : submodule R (dual R M)`, + pulled back along `module.dual.eval R M`. + * `submodule.dual_copairing W` is the canonical pairing between `W.dual_annihilator` and `M ⧸ W`. + It is nondegenerate for vector spaces (`subspace.dual_copairing_nondegenerate`). + * `submodule.dual_pairing W` is the canonical pairing between `dual R M ⧸ W.dual_annihilator` + and `W`. It is nondegenerate for vector spaces (`subspace.dual_pairing_nondegenerate`). +* Vector spaces: + * `subspace.dual_lift W` is an arbitrary section (using choice) of `submodule.dual_restrict W`. ## Main results -* `to_dual_equiv` : the linear equivalence between the dual module and primal module, - given a finite basis. -* `module.dual_bases.basis` and `module.dual_bases.eq_dual`: if `e` and `ε` form a dual pair, `e` - is a basis and `ε` is its dual basis. -* `quot_equiv_annihilator`: the quotient by a subspace is isomorphic to its dual annihilator. - -## Notation - -We sometimes use `V'` as local notation for `dual K V`. +* Bases: + * `module.dual_basis.basis` and `module.dual_basis.coe_basis`: if `e` and `ε` form a dual pair, + then `e` is a basis. + * `module.dual_basis.coe_dual_basis`: if `e` and `ε` form a dual pair, + then `ε` is a basis. +* Annihilators: + * `module.dual_annihilator_gc R M` is the antitone Galois correspondence between + `submodule.dual_annihilator` and `submodule.dual_coannihilator`. + * `linear_map.ker_dual_map_eq_dual_annihilator_range` says that + `f.dual_map.ker = f.range.dual_annihilator` + * `linear_map.range_dual_map_eq_dual_annihilator_ker_of_subtype_range_surjective` says that + `f.dual_map.range = f.ker.dual_annihilator`; this is specialized to vector spaces in + `linear_map.range_dual_map_eq_dual_annihilator_ker`. + * `submodule.dual_quot_equiv_dual_annihilator` is the equivalence + `dual R (M ⧸ W) ≃ₗ[R] W.dual_annihilator` +* Vector spaces: + * `subspace.dual_annihilator_dual_coannihilator_eq` says that the double dual annihilator, + pulled back ground `module.dual.eval`, is the original submodule. + * `subspace.dual_annihilator_gci` says that `module.dual_annihilator_gc R M` is an + antitone Galois coinsertion. + * `subspace.quot_annihilator_equiv` is the equivalence + `dual K V ⧸ W.dual_annihilator ≃ₗ[K] dual K W`. + * `linear_map.dual_pairing_nondegenerate` says that `module.dual_pairing` is nondegenerate. + * `subspace.is_compl_dual_annihilator` says that the dual annihilator carries complementary + subspaces to complementary subspaces. +* Finite-dimensional vector spaces: + * `module.eval_equiv` is the equivalence `V ≃ₗ[K] dual K (dual K V)` + * `module.map_eval_equiv` is the order isomorphism between subspaces of `V` and + subspaces of `dual K (dual K V)`. + * `subspace.quot_dual_equiv_annihilator W` is the equivalence + `(dual K V ⧸ W.dual_lift.range) ≃ₗ[K] W.dual_annihilator`, where `W.dual_lift.range` is a copy + of `dual K W` inside `dual K V`. + * `subspace.quot_equiv_annihilator W` is the equivalence `(V ⧸ W) ≃ₗ[K] W.dual_annihilator` + * `subspace.dual_quot_distrib W` is an equivalence + `dual K (V₁ ⧸ W) ≃ₗ[K] dual K V₁ ⧸ W.dual_lift.range` from an arbitrary choice of + splitting of `V₁`. ## TODO @@ -93,8 +140,93 @@ lemma transpose_comp (u : M' →ₗ[R] M'') (v : M →ₗ[R] M') : end dual +section prod +variables (M' : Type*) [add_comm_monoid M'] [module R M'] + +/-- Taking duals distributes over products. -/ +@[simps] def dual_prod_dual_equiv_dual : + (module.dual R M × module.dual R M') ≃ₗ[R] module.dual R (M × M') := +linear_map.coprod_equiv R + +@[simp] lemma dual_prod_dual_equiv_dual_apply (φ : module.dual R M) (ψ : module.dual R M') : + dual_prod_dual_equiv_dual R M M' (φ, ψ) = φ.coprod ψ := rfl + +end prod + end module +section dual_map +open module + +variables {R : Type*} [comm_semiring R] {M₁ : Type*} {M₂ : Type*} +variables [add_comm_monoid M₁] [module R M₁] [add_comm_monoid M₂] [module R M₂] + +/-- Given a linear map `f : M₁ →ₗ[R] M₂`, `f.dual_map` is the linear map between the dual of +`M₂` and `M₁` such that it maps the functional `φ` to `φ ∘ f`. -/ +def linear_map.dual_map (f : M₁ →ₗ[R] M₂) : dual R M₂ →ₗ[R] dual R M₁ := +module.dual.transpose f + +lemma linear_map.dual_map_def (f : M₁ →ₗ[R] M₂) : f.dual_map = module.dual.transpose f := rfl + +lemma linear_map.dual_map_apply' (f : M₁ →ₗ[R] M₂) (g : dual R M₂) : + f.dual_map g = g.comp f := rfl + +@[simp] lemma linear_map.dual_map_apply (f : M₁ →ₗ[R] M₂) (g : dual R M₂) (x : M₁) : + f.dual_map g x = g (f x) := rfl + +@[simp] lemma linear_map.dual_map_id : + (linear_map.id : M₁ →ₗ[R] M₁).dual_map = linear_map.id := +by { ext, refl } + +lemma linear_map.dual_map_comp_dual_map {M₃ : Type*} [add_comm_group M₃] [module R M₃] + (f : M₁ →ₗ[R] M₂) (g : M₂ →ₗ[R] M₃) : + f.dual_map.comp g.dual_map = (g.comp f).dual_map := +rfl + +/-- If a linear map is surjective, then its dual is injective. -/ +lemma linear_map.dual_map_injective_of_surjective {f : M₁ →ₗ[R] M₂} (hf : function.surjective f) : + function.injective f.dual_map := +begin + intros φ ψ h, + ext x, + obtain ⟨y, rfl⟩ := hf x, + exact congr_arg (λ (g : module.dual R M₁), g y) h, +end + +/-- The `linear_equiv` version of `linear_map.dual_map`. -/ +def linear_equiv.dual_map (f : M₁ ≃ₗ[R] M₂) : dual R M₂ ≃ₗ[R] dual R M₁ := +{ inv_fun := f.symm.to_linear_map.dual_map, + left_inv := + begin + intro φ, ext x, + simp only [linear_map.dual_map_apply, linear_equiv.coe_to_linear_map, + linear_map.to_fun_eq_coe, linear_equiv.apply_symm_apply] + end, + right_inv := + begin + intro φ, ext x, + simp only [linear_map.dual_map_apply, linear_equiv.coe_to_linear_map, + linear_map.to_fun_eq_coe, linear_equiv.symm_apply_apply] + end, + .. f.to_linear_map.dual_map } + +@[simp] lemma linear_equiv.dual_map_apply (f : M₁ ≃ₗ[R] M₂) (g : dual R M₂) (x : M₁) : + f.dual_map g x = g (f x) := rfl + +@[simp] lemma linear_equiv.dual_map_refl : + (linear_equiv.refl R M₁).dual_map = linear_equiv.refl R (dual R M₁) := +by { ext, refl } + +@[simp] lemma linear_equiv.dual_map_symm {f : M₁ ≃ₗ[R] M₂} : + (linear_equiv.dual_map f).symm = linear_equiv.dual_map f.symm := rfl + +lemma linear_equiv.dual_map_trans {M₃ : Type*} [add_comm_group M₃] [module R M₃] + (f : M₁ ≃ₗ[R] M₂) (g : M₂ ≃ₗ[R] M₃) : + g.dual_map.trans f.dual_map = (f.trans g).dual_map := +rfl + +end dual_map + namespace basis universes u v w @@ -314,14 +446,33 @@ variables {K V : Type*} variables [field K] [add_comm_group V] [module K V] open module module.dual submodule linear_map cardinal basis finite_dimensional +section +variables (K) (V) + theorem eval_ker : (eval K V).ker = ⊥ := by { classical, exact (basis.of_vector_space K V).eval_ker } +theorem map_eval_injective : (submodule.map (eval K V)).injective := +begin + apply submodule.map_injective_of_injective, + rw ← linear_map.ker_eq_bot, + apply eval_ker K V, -- elaborates faster than `exact` +end + +theorem comap_eval_surjective : (submodule.comap (eval K V)).surjective := +begin + apply submodule.comap_surjective_of_injective, + rw ← linear_map.ker_eq_bot, + apply eval_ker K V, -- elaborates faster than `exact` +end + +end + section variable (K) theorem eval_apply_eq_zero_iff (v : V) : (eval K V) v = 0 ↔ v = 0 := -by simpa only using set_like.ext_iff.mp (eval_ker : (eval K V).ker = _) v +by simpa only using set_like.ext_iff.mp (eval_ker K V) v theorem eval_apply_injective : function.injective (eval K V) := (injective_iff_map_eq_zero' (eval K V)).mpr (eval_apply_eq_zero_iff K) @@ -347,13 +498,25 @@ variables (K V) /-- A vector space is linearly equivalent to the dual of its dual space. -/ def eval_equiv [finite_dimensional K V] : V ≃ₗ[K] dual K (dual K V) := linear_equiv.of_bijective (eval K V) - ⟨ker_eq_bot.mp eval_ker, range_eq_top.mp erange_coe⟩ + -- 60x faster elaboration than using `ker_eq_bot.mp eval_ker` directly: + ⟨by { rw ← ker_eq_bot, apply eval_ker K V }, range_eq_top.mp erange_coe⟩ + +/-- The isomorphism `module.eval_equiv` induces an order isomorphism on subspaces. -/ +def map_eval_equiv [finite_dimensional K V] : subspace K V ≃o subspace K (dual K (dual K V)) := +submodule.order_iso_map_comap (eval_equiv K V) variables {K V} @[simp] lemma eval_equiv_to_linear_map [finite_dimensional K V] : (eval_equiv K V).to_linear_map = dual.eval K V := rfl +@[simp] lemma map_eval_equiv_apply [finite_dimensional K V] (W : subspace K V) : + map_eval_equiv K V W = W.map (eval K V) := rfl + +@[simp] lemma map_eval_equiv_symm_apply [finite_dimensional K V] + (W'' : subspace K (dual K (dual K V))) : + (map_eval_equiv K V).symm W'' = W''.comap (eval K V) := rfl + end module section dual_bases @@ -363,12 +526,17 @@ open module variables {R M ι : Type*} variables [comm_semiring R] [add_comm_monoid M] [module R M] [decidable_eq ι] +/-- Try using `set.to_finite` to dispatch a `set.finite` goal. -/ +-- TODO: In Lean 4 we can remove this and use `by { intros; exact Set.toFinite _ }` as a default +-- argument. +meta def use_finite_instance : tactic unit := `[intros, exact set.to_finite _] + /-- `e` and `ε` have characteristic properties of a basis and its dual -/ @[nolint has_nonempty_instance] -structure module.dual_bases (e : ι → M) (ε : ι → (dual R M)) := +structure module.dual_bases (e : ι → M) (ε : ι → (dual R M)) : Prop := (eval : ∀ i j : ι, ε i (e j) = if i = j then 1 else 0) (total : ∀ {m : M}, (∀ i, ε i m = 0) → m = 0) -[finite : ∀ m : M, fintype {i | ε i m ≠ 0}] +(finite : ∀ m : M, {i | ε i m ≠ 0}.finite . use_finite_instance) end dual_bases @@ -383,8 +551,8 @@ variables {e : ι → M} {ε : ι → dual R M} /-- The coefficients of `v` on the basis `e` -/ def coeffs [decidable_eq ι] (h : dual_bases e ε) (m : M) : ι →₀ R := { to_fun := λ i, ε i m, - support := by { haveI := h.finite m, exact {i : ι | ε i m ≠ 0}.to_finset }, - mem_support_to_fun := by {intro i, rw set.mem_to_finset, exact iff.rfl } } + support := (h.finite m).to_finset, + mem_support_to_fun := by { intro i, rw [set.finite.mem_to_finset, set.mem_set_of_eq] } } @[simp] lemma coeffs_apply [decidable_eq ι] (h : dual_bases e ε) (m : M) (i : ι) : h.coeffs m i = ε i m := rfl @@ -470,6 +638,8 @@ def dual_restrict (W : submodule R M) : module.dual R M →ₗ[R] module.dual R W := linear_map.dom_restrict' W +lemma dual_restrict_def (W : submodule R M) : W.dual_restrict = W.subtype.dual_map := rfl + @[simp] lemma dual_restrict_apply (W : submodule R M) (φ : module.dual R M) (x : W) : W.dual_restrict φ x = φ (x : M) := rfl @@ -488,136 +658,115 @@ begin exact ⟨λ h w hw, h ⟨w, hw⟩, λ h w, h w.1 w.2⟩ end +/-- That $\operatorname{ker}(\iota^* : V^* \to W^*) = \operatorname{ann}(W)$. +This is the definition of the dual annihilator of the submodule $W$. -/ lemma dual_restrict_ker_eq_dual_annihilator (W : submodule R M) : W.dual_restrict.ker = W.dual_annihilator := rfl /-- The `dual_annihilator` of a submodule of the dual space pulled back along the evaluation map `module.dual.eval`. -/ -def dual_annihilator_comap (Φ : submodule R (module.dual R M)) : submodule R M := +def dual_coannihilator (Φ : submodule R (module.dual R M)) : submodule R M := Φ.dual_annihilator.comap (module.dual.eval R M) -lemma mem_dual_annihilator_comap {Φ : submodule R (module.dual R M)} (x : M) : - x ∈ Φ.dual_annihilator_comap ↔ ∀ φ ∈ Φ, (φ x : R) = 0 := -by simp_rw [dual_annihilator_comap, mem_comap, mem_dual_annihilator, module.dual.eval_apply] +lemma mem_dual_coannihilator {Φ : submodule R (module.dual R M)} (x : M) : + x ∈ Φ.dual_coannihilator ↔ ∀ φ ∈ Φ, (φ x : R) = 0 := +by simp_rw [dual_coannihilator, mem_comap, mem_dual_annihilator, module.dual.eval_apply] -@[simp] lemma dual_annihilator_top : (⊤ : submodule R M).dual_annihilator = ⊥ := +lemma dual_annihilator_gc (R M : Type*) [comm_semiring R] [add_comm_monoid M] [module R M] : + galois_connection + (order_dual.to_dual ∘ (dual_annihilator : submodule R M → submodule R (module.dual R M))) + (dual_coannihilator ∘ order_dual.of_dual) := begin - rw eq_bot_iff, - intro v, - simp_rw [mem_dual_annihilator, mem_bot, mem_top, forall_true_left], - exact λ h, linear_map.ext h, + intros a b, + induction b using order_dual.rec, + simp only [function.comp_app, order_dual.to_dual_le_to_dual, order_dual.of_dual_to_dual], + split; + { intros h x hx, + simp only [mem_dual_annihilator, mem_dual_coannihilator], + intros y hy, + have := h hy, + simp only [mem_dual_annihilator, mem_dual_coannihilator] at this, + exact this x hx }, end +lemma le_dual_annihilator_iff_le_dual_coannihilator + {U : submodule R (module.dual R M)} {V : submodule R M} : + U ≤ V.dual_annihilator ↔ V ≤ U.dual_coannihilator := +(dual_annihilator_gc R M).le_iff_le + @[simp] lemma dual_annihilator_bot : (⊥ : submodule R M).dual_annihilator = ⊤ := +(dual_annihilator_gc R M).l_bot + +@[simp] lemma dual_annihilator_top : (⊤ : submodule R M).dual_annihilator = ⊥ := begin - rw eq_top_iff, + rw eq_bot_iff, intro v, simp_rw [mem_dual_annihilator, mem_bot, mem_top, forall_true_left], - rintro _ rfl, - exact _root_.map_zero v, + exact λ h, linear_map.ext h, end -@[simp] lemma dual_annihilator_comap_bot : - (⊥ : submodule R (module.dual R M)).dual_annihilator_comap = ⊤ := -by rw [dual_annihilator_comap, dual_annihilator_bot, comap_top] +@[simp] lemma dual_coannihilator_bot : + (⊥ : submodule R (module.dual R M)).dual_coannihilator = ⊤ := +(dual_annihilator_gc R M).u_top @[mono] lemma dual_annihilator_anti {U V : submodule R M} (hUV : U ≤ V) : V.dual_annihilator ≤ U.dual_annihilator := -begin - intro φ, - simp_rw [mem_dual_annihilator], - intros h w hw, - exact h w (hUV hw), -end +(dual_annihilator_gc R M).monotone_l hUV -@[mono] lemma dual_annihilator_comap_anti {U V : submodule R (module.dual R M)} (hUV : U ≤ V) : - V.dual_annihilator_comap ≤ U.dual_annihilator_comap := -begin - intro φ, - simp_rw [mem_dual_annihilator_comap], - intros h w hw, - exact h w (hUV hw), -end +@[mono] lemma dual_coannihilator_anti {U V : submodule R (module.dual R M)} (hUV : U ≤ V) : + V.dual_coannihilator ≤ U.dual_coannihilator := +(dual_annihilator_gc R M).monotone_u hUV -lemma le_dual_annihilator_dual_annihilator_comap {U : submodule R M} : - U ≤ U.dual_annihilator.dual_annihilator_comap := -begin - intro v, - simp_rw [mem_dual_annihilator_comap, mem_dual_annihilator], - intros hv φ h, - exact h _ hv, -end +lemma le_dual_annihilator_dual_coannihilator (U : submodule R M) : + U ≤ U.dual_annihilator.dual_coannihilator := +(dual_annihilator_gc R M).le_u_l U -lemma le_dual_annihilator_comap_dual_annihilator {U : submodule R (module.dual R M)} : - U ≤ U.dual_annihilator_comap.dual_annihilator := -begin - intro v, - simp_rw [mem_dual_annihilator, mem_dual_annihilator_comap], - intros hv φ h, - exact h _ hv, -end +lemma le_dual_coannihilator_dual_annihilator (U : submodule R (module.dual R M)) : + U ≤ U.dual_coannihilator.dual_annihilator := +(dual_annihilator_gc R M).l_u_le U + +lemma dual_annihilator_dual_coannihilator_dual_annihilator + (U : submodule R M) : + U.dual_annihilator.dual_coannihilator.dual_annihilator = U.dual_annihilator := +(dual_annihilator_gc R M).l_u_l_eq_l U + +lemma dual_coannihilator_dual_annihilator_dual_coannihilator + (U : submodule R (module.dual R M)) : + U.dual_coannihilator.dual_annihilator.dual_coannihilator = U.dual_coannihilator := +(dual_annihilator_gc R M).u_l_u_eq_u U lemma dual_annihilator_sup_eq (U V : submodule R M) : (U ⊔ V).dual_annihilator = U.dual_annihilator ⊓ V.dual_annihilator := -begin - ext φ, - rw [mem_inf, mem_dual_annihilator, mem_dual_annihilator, mem_dual_annihilator], - split; intro h, - { refine ⟨_, _⟩; - intros x hx, - exact h x (mem_sup.2 ⟨x, hx, 0, zero_mem _, add_zero _⟩), - exact h x (mem_sup.2 ⟨0, zero_mem _, x, hx, zero_add _⟩) }, - { simp_rw mem_sup, - rintro _ ⟨x, hx, y, hy, rfl⟩, - rw [linear_map.map_add, h.1 _ hx, h.2 _ hy, add_zero] } -end +(dual_annihilator_gc R M).l_sup + +lemma dual_coannihilator_sup_eq (U V : submodule R (module.dual R M)) : + (U ⊔ V).dual_coannihilator = U.dual_coannihilator ⊓ V.dual_coannihilator := +(dual_annihilator_gc R M).u_inf lemma dual_annihilator_supr_eq {ι : Type*} (U : ι → submodule R M) : (⨆ (i : ι), U i).dual_annihilator = ⨅ (i : ι), (U i).dual_annihilator := -begin - classical, - ext φ, - simp_rw [mem_infi, mem_dual_annihilator], - split, - { simp_rw [mem_supr], - intros h i w hw, - exact h _ (λ _ hi, hi i hw), }, - { simp_rw [submodule.mem_supr_iff_exists_dfinsupp'], - rintros h w ⟨f, rfl⟩, - simp only [linear_map.map_dfinsupp_sum], - transitivity f.sum (λ (i : ι) (d : U i), (0 : R)), - { congr, - ext i d, - exact h i d d.property, }, - { exact @dfinsupp.sum_zero ι _ (λ i, U i) _ _ _ _ f, } } -end +(dual_annihilator_gc R M).l_supr + +lemma dual_coannihilator_supr_eq {ι : Type*} (U : ι → submodule R (module.dual R M)) : + (⨆ (i : ι), U i).dual_coannihilator = ⨅ (i : ι), (U i).dual_coannihilator := +(dual_annihilator_gc R M).u_infi --- TODO: when `M` is finite-dimensional this is an equality +/-- See also `subspace.dual_annihilator_inf_eq` for vector subspaces. -/ lemma sup_dual_annihilator_le_inf (U V : submodule R M) : U.dual_annihilator ⊔ V.dual_annihilator ≤ (U ⊓ V).dual_annihilator := begin - intro φ, - simp_rw [mem_sup, mem_dual_annihilator, mem_inf], - rintro ⟨ψ, hψ, ψ', hψ', rfl⟩ v ⟨hU, hV⟩, - rw [linear_map.add_apply, hψ _ hU, hψ' _ hV, zero_add], + rw [le_dual_annihilator_iff_le_dual_coannihilator, dual_coannihilator_sup_eq], + apply' inf_le_inf; exact le_dual_annihilator_dual_coannihilator _, end --- TODO: when `M` is finite-dimensional this is an equality +/-- See also `subspace.dual_annihilator_infi_eq` for vector subspaces when `ι` is finite. -/ lemma supr_dual_annihilator_le_infi {ι : Type*} (U : ι → submodule R M) : (⨆ (i : ι), (U i).dual_annihilator) ≤ (⨅ (i : ι), U i).dual_annihilator := begin - classical, - intro φ, - simp_rw [mem_dual_annihilator, submodule.mem_supr_iff_exists_dfinsupp', mem_infi], - rintros ⟨f, rfl⟩ x hx, - rw [linear_map.dfinsupp_sum_apply], - transitivity f.sum (λ (i : ι) (d : (U i).dual_annihilator), (0 : R)), - { congr, - ext i ⟨d, hd⟩, - rw [mem_dual_annihilator] at hd, - exact hd x (hx _), }, - { exact @dfinsupp.sum_zero ι _ (λ i, (U i).dual_annihilator) _ _ _ _ f } + rw [le_dual_annihilator_iff_le_dual_coannihilator, dual_coannihilator_supr_eq], + apply' infi_mono, + exact λ (i : ι), le_dual_annihilator_dual_coannihilator (U i), end end submodule @@ -631,16 +780,16 @@ universes u v w -- We work in vector spaces because `exists_is_compl` only hold for vector spaces variables {K : Type u} {V : Type v} [field K] [add_comm_group V] [module K V] -@[simp] lemma dual_annihilator_comap_top (W : subspace K V) : - (⊤ : submodule K (module.dual K W)).dual_annihilator_comap = ⊥ := -by rw [dual_annihilator_comap, dual_annihilator_top, comap_bot, module.eval_ker] +@[simp] lemma dual_coannihilator_top (W : subspace K V) : + (⊤ : subspace K (module.dual K W)).dual_coannihilator = ⊥ := +by rw [dual_coannihilator, dual_annihilator_top, comap_bot, module.eval_ker] -lemma dual_annihilator_dual_annihilator_comap_eq {W : subspace K V} : - W.dual_annihilator.dual_annihilator_comap = W := +lemma dual_annihilator_dual_coannihilator_eq {W : subspace K V} : + W.dual_annihilator.dual_coannihilator = W := begin - refine le_antisymm _ le_dual_annihilator_dual_annihilator_comap, + refine le_antisymm _ (le_dual_annihilator_dual_coannihilator _), intro v, - simp only [mem_dual_annihilator, mem_dual_annihilator_comap], + simp only [mem_dual_annihilator, mem_dual_coannihilator], contrapose!, intro hv, obtain ⟨W', hW⟩ := submodule.exists_is_compl W, @@ -655,12 +804,39 @@ begin of_is_compl_left_apply, zero_apply, of_is_compl_right_apply, zero_add, ne.def], refine ⟨_, hφ⟩, intros v hv, - convert linear_map.of_is_compl_left_apply hW ⟨v, hv⟩, + apply linear_map.of_is_compl_left_apply hW ⟨v, hv⟩, -- exact elaborates slowly +end + +theorem forall_mem_dual_annihilator_apply_eq_zero_iff (W : subspace K V) (v : V) : + (∀ (φ : module.dual K V), φ ∈ W.dual_annihilator → φ v = 0) ↔ v ∈ W := +by rw [← set_like.ext_iff.mp dual_annihilator_dual_coannihilator_eq v, + mem_dual_coannihilator] + +/-- `submodule.dual_annihilator` and `submodule.dual_coannihilator` form a Galois coinsertion. -/ +def dual_annihilator_gci (K V : Type*) [field K] [add_comm_group V] [module K V] : + galois_coinsertion + (order_dual.to_dual ∘ (dual_annihilator : subspace K V → subspace K (module.dual K V))) + (dual_coannihilator ∘ order_dual.of_dual) := +{ choice := λ W h, dual_coannihilator W, + gc := dual_annihilator_gc K V, + u_l_le := λ W, dual_annihilator_dual_coannihilator_eq.le, + choice_eq := λ W h, rfl } + +lemma dual_annihilator_le_dual_annihilator_iff {W W' : subspace K V} : + W.dual_annihilator ≤ W'.dual_annihilator ↔ W' ≤ W := +(dual_annihilator_gci K V).l_le_l_iff + +lemma dual_annihilator_inj {W W' : subspace K V} : + W.dual_annihilator = W'.dual_annihilator ↔ W = W' := +begin + split, + { apply (dual_annihilator_gci K V).l_injective }, + { rintro rfl, refl }, end /-- Given a subspace `W` of `V` and an element of its dual `φ`, `dual_lift W φ` is -the natural extension of `φ` to an element of the dual of `V`. -That is, `dual_lift W φ` sends `w ∈ W` to `φ x` and `x` in the complement of `W` to `0`. -/ +an arbitrary extension of `φ` to an element of the dual of `V`. +That is, `dual_lift W φ` sends `w ∈ W` to `φ x` and `x` in a chosen complement of `W` to `0`. -/ noncomputable def dual_lift (W : subspace K V) : module.dual K W →ₗ[K] module.dual K V := let h := classical.indefinite_description _ W.exists_is_compl in @@ -703,7 +879,11 @@ noncomputable def quot_annihilator_equiv (W : subspace K V) : (quot_equiv_of_eq _ _ W.dual_restrict_ker_eq_dual_annihilator).symm.trans $ W.dual_restrict.quot_ker_equiv_of_surjective dual_restrict_surjective -/-- The natural isomorphism forom the dual of a subspace `W` to `W.dual_lift.range`. -/ +@[simp] lemma quot_annihilator_equiv_apply (W : subspace K V) (φ : module.dual K V) : + W.quot_annihilator_equiv (submodule.quotient.mk φ) = W.dual_restrict φ := +by { ext, refl } + +/-- The natural isomorphism from the dual of a subspace `W` to `W.dual_lift.range`. -/ noncomputable def dual_equiv_dual (W : subspace K V) : module.dual K W ≃ₗ[K] W.dual_lift.range := linear_equiv.of_injective _ dual_lift_injective @@ -727,6 +907,15 @@ by apply_instance variables [finite_dimensional K V] [finite_dimensional K V₁] +lemma dual_annihilator_dual_annihilator_eq (W : subspace K V) : + W.dual_annihilator.dual_annihilator = module.map_eval_equiv K V W := +begin + have : _ = W := subspace.dual_annihilator_dual_coannihilator_eq, + rw [dual_coannihilator, ← module.map_eval_equiv_symm_apply] at this, + rwa ← order_iso.symm_apply_eq, +end + +-- TODO(kmill): https://github.com/leanprover-community/mathlib/pull/17521#discussion_r1083241963 @[simp] lemma dual_finrank_eq : finrank K (module.dual K V) = finrank K V := linear_equiv.finrank_eq (basis.of_vector_space K V).to_dual_equiv.symm @@ -749,18 +938,18 @@ end open finite_dimensional @[simp] -lemma finrank_dual_annihilator_comap_eq {Φ : subspace K (module.dual K V)} : - finrank K Φ.dual_annihilator_comap = finrank K Φ.dual_annihilator := +lemma finrank_dual_coannihilator_eq {Φ : subspace K (module.dual K V)} : + finrank K Φ.dual_coannihilator = finrank K Φ.dual_annihilator := begin - rw [submodule.dual_annihilator_comap, ← module.eval_equiv_to_linear_map], + rw [submodule.dual_coannihilator, ← module.eval_equiv_to_linear_map], exact linear_equiv.finrank_eq (linear_equiv.of_submodule' _ _), end -lemma finrank_add_finrank_dual_annihilator_comap_eq +lemma finrank_add_finrank_dual_coannihilator_eq (W : subspace K (module.dual K V)) : - finrank K W + finrank K W.dual_annihilator_comap = finrank K V := + finrank K W + finrank K W.dual_coannihilator = finrank K V := begin - rw [finrank_dual_annihilator_comap_eq, W.quot_equiv_annihilator.finrank_eq.symm, add_comm, + rw [finrank_dual_coannihilator_eq, W.quot_equiv_annihilator.finrank_eq.symm, add_comm, submodule.finrank_quotient_add_finrank, subspace.dual_finrank_eq], end @@ -770,61 +959,6 @@ end subspace open module -section dual_map -variables {R : Type*} [comm_semiring R] {M₁ : Type*} {M₂ : Type*} -variables [add_comm_monoid M₁] [module R M₁] [add_comm_monoid M₂] [module R M₂] - -/-- Given a linear map `f : M₁ →ₗ[R] M₂`, `f.dual_map` is the linear map between the dual of -`M₂` and `M₁` such that it maps the functional `φ` to `φ ∘ f`. -/ -def linear_map.dual_map (f : M₁ →ₗ[R] M₂) : dual R M₂ →ₗ[R] dual R M₁ := -linear_map.lcomp R R f - -@[simp] lemma linear_map.dual_map_apply (f : M₁ →ₗ[R] M₂) (g : dual R M₂) (x : M₁) : - f.dual_map g x = g (f x) := rfl - -@[simp] lemma linear_map.dual_map_id : - (linear_map.id : M₁ →ₗ[R] M₁).dual_map = linear_map.id := -by { ext, refl } - -lemma linear_map.dual_map_comp_dual_map {M₃ : Type*} [add_comm_group M₃] [module R M₃] - (f : M₁ →ₗ[R] M₂) (g : M₂ →ₗ[R] M₃) : - f.dual_map.comp g.dual_map = (g.comp f).dual_map := -rfl - -/-- The `linear_equiv` version of `linear_map.dual_map`. -/ -def linear_equiv.dual_map (f : M₁ ≃ₗ[R] M₂) : dual R M₂ ≃ₗ[R] dual R M₁ := -{ inv_fun := f.symm.to_linear_map.dual_map, - left_inv := - begin - intro φ, ext x, - simp only [linear_map.dual_map_apply, linear_equiv.coe_to_linear_map, - linear_map.to_fun_eq_coe, linear_equiv.apply_symm_apply] - end, - right_inv := - begin - intro φ, ext x, - simp only [linear_map.dual_map_apply, linear_equiv.coe_to_linear_map, - linear_map.to_fun_eq_coe, linear_equiv.symm_apply_apply] - end, - .. f.to_linear_map.dual_map } - -@[simp] lemma linear_equiv.dual_map_apply (f : M₁ ≃ₗ[R] M₂) (g : dual R M₂) (x : M₁) : - f.dual_map g x = g (f x) := rfl - -@[simp] lemma linear_equiv.dual_map_refl : - (linear_equiv.refl R M₁).dual_map = linear_equiv.refl R (dual R M₁) := -by { ext, refl } - -@[simp] lemma linear_equiv.dual_map_symm {f : M₁ ≃ₗ[R] M₂} : - (linear_equiv.dual_map f).symm = linear_equiv.dual_map f.symm := rfl - -lemma linear_equiv.dual_map_trans {M₃ : Type*} [add_comm_group M₃] [module R M₃] - (f : M₁ ≃ₗ[R] M₂) (g : M₂ ≃ₗ[R] M₃) : - g.dual_map.trans f.dual_map = (f.trans g).dual_map := -rfl - -end dual_map - namespace linear_map variables {R : Type*} [comm_semiring R] {M₁ : Type*} {M₂ : Type*} variables [add_comm_monoid M₁] [module R M₁] [add_comm_monoid M₂] [module R M₂] @@ -854,25 +988,257 @@ begin rw [dual_map_apply, hx, map_zero] end -/-- if a linear map is surjective, then its dual is injective -/ -lemma dual_map_injective_of_surjective {f : M₁ →ₗ[R] M₂} (hf : function.surjective f) : - function.injective f.dual_map := +end linear_map + +section comm_ring + +variables {R M M' : Type*} +variables [comm_ring R] [add_comm_group M] [module R M] [add_comm_group M'] [module R M'] + +namespace submodule + +/-- Given a submodule, corestrict to the pairing on `M ⧸ W` by +simultaneously restricting to `W.dual_annihilator`. + +See `subspace.dual_copairing_nondegenerate`. -/ +def dual_copairing (W : submodule R M) : + W.dual_annihilator →ₗ[R] M ⧸ W →ₗ[R] R := +linear_map.flip $ W.liftq ((module.dual_pairing R M).dom_restrict W.dual_annihilator).flip + (by { intros w hw, ext ⟨φ, hφ⟩, exact (mem_dual_annihilator φ).mp hφ w hw }) + +@[simp] lemma dual_copairing_apply {W : submodule R M} (φ : W.dual_annihilator) (x : M) : + W.dual_copairing φ (quotient.mk x) = φ x := rfl + +/-- Given a submodule, restrict to the pairing on `W` by +simultaneously corestricting to `module.dual R M ⧸ W.dual_annihilator`. +This is `submodule.dual_restrict` factored through the quotient by its kernel (which +is `W.dual_annihilator` by definition). + +See `subspace.dual_pairing_nondegenerate`. -/ +def dual_pairing (W : submodule R M) : + module.dual R M ⧸ W.dual_annihilator →ₗ[R] W →ₗ[R] R := +W.dual_annihilator.liftq W.dual_restrict le_rfl + +@[simp] lemma dual_pairing_apply {W : submodule R M} (φ : module.dual R M) (x : W) : + W.dual_pairing (quotient.mk φ) x = φ x := rfl + +/-- That $\operatorname{im}(q^* : (V/W)^* \to V^*) = \operatorname{ann}(W)$. -/ +lemma range_dual_map_mkq_eq (W : submodule R M) : + W.mkq.dual_map.range = W.dual_annihilator := begin - obtain ⟨g, hg⟩ := hf.has_right_inverse, - intros x y hxy, - ext w, - simpa only [dual_map_apply, hg w] using fun_like.congr_fun hxy (g w), + ext φ, + rw linear_map.mem_range, + split, + { rintro ⟨ψ, rfl⟩, + have := linear_map.mem_range_self W.mkq.dual_map ψ, + simpa only [ker_mkq] using linear_map.range_dual_map_le_dual_annihilator_ker W.mkq this, }, + { intro hφ, + existsi W.dual_copairing ⟨φ, hφ⟩, + ext, + refl, } end -section finite_dimensional +/-- Equivalence $(M/W)^* \approx \operatorname{ann}(W)$. That is, there is a one-to-one +correspondence between the dual of `M ⧸ W` and those elements of the dual of `M` that +vanish on `W`. + +The inverse of this is `submodule.dual_copairing`. -/ +def dual_quot_equiv_dual_annihilator (W : submodule R M) : + module.dual R (M ⧸ W) ≃ₗ[R] W.dual_annihilator := +linear_equiv.of_linear + (W.mkq.dual_map.cod_restrict W.dual_annihilator $ + λ φ, W.range_dual_map_mkq_eq ▸ W.mkq.dual_map.mem_range_self φ) + W.dual_copairing + (by { ext, refl}) (by { ext, refl }) + +@[simp] lemma dual_quot_equiv_dual_annihilator_apply (W : submodule R M) + (φ : module.dual R (M ⧸ W)) (x : M) : + dual_quot_equiv_dual_annihilator W φ x = φ (quotient.mk x) := rfl + +lemma dual_copairing_eq (W : submodule R M) : + W.dual_copairing = (dual_quot_equiv_dual_annihilator W).symm.to_linear_map := rfl + +@[simp] lemma dual_quot_equiv_dual_annihilator_symm_apply_mk (W : submodule R M) + (φ : W.dual_annihilator) (x : M) : + (dual_quot_equiv_dual_annihilator W).symm φ (quotient.mk x) = φ x := rfl + +end submodule + +namespace linear_map +open submodule + +lemma range_dual_map_eq_dual_annihilator_ker_of_surjective + (f : M →ₗ[R] M') (hf : function.surjective f) : + f.dual_map.range = f.ker.dual_annihilator := +begin + rw ← f.ker.range_dual_map_mkq_eq, + let f' := linear_map.quot_ker_equiv_of_surjective f hf, + transitivity linear_map.range (f.dual_map.comp f'.symm.dual_map.to_linear_map), + { rw linear_map.range_comp_of_range_eq_top, + apply linear_equiv.range }, + { apply congr_arg, + ext φ x, + simp only [linear_map.coe_comp, linear_equiv.coe_to_linear_map, linear_map.dual_map_apply, + linear_equiv.dual_map_apply, mkq_apply, f', linear_map.quot_ker_equiv_of_surjective, + linear_equiv.trans_symm, linear_equiv.trans_apply, linear_equiv.of_top_symm_apply, + linear_map.quot_ker_equiv_range_symm_apply_image, mkq_apply], } +end + +-- Note, this can be specialized to the case where `R` is an injective `R`-module, or when +-- `f.coker` is a projective `R`-module. +lemma range_dual_map_eq_dual_annihilator_ker_of_subtype_range_surjective + (f : M →ₗ[R] M') (hf : function.surjective f.range.subtype.dual_map) : + f.dual_map.range = f.ker.dual_annihilator := +begin + have rr_surj : function.surjective f.range_restrict, + { rw [← linear_map.range_eq_top, linear_map.range_range_restrict] }, + have := range_dual_map_eq_dual_annihilator_ker_of_surjective f.range_restrict rr_surj, + convert this using 1, + { change ((submodule.subtype f.range).comp f.range_restrict).dual_map.range = _, + rw [← linear_map.dual_map_comp_dual_map, linear_map.range_comp_of_range_eq_top], + rwa linear_map.range_eq_top, }, + { apply congr_arg, + exact (linear_map.ker_range_restrict f).symm, }, +end + +end linear_map + +end comm_ring + +section vector_space variables {K : Type*} [field K] {V₁ : Type*} {V₂ : Type*} variables [add_comm_group V₁] [module K V₁] [add_comm_group V₂] [module K V₂] -open finite_dimensional +namespace linear_map + +lemma dual_pairing_nondegenerate : (dual_pairing K V₁).nondegenerate := +⟨separating_left_iff_ker_eq_bot.mpr ker_id, λ x, (forall_dual_apply_eq_zero_iff K x).mp⟩ + +lemma dual_map_surjective_of_injective {f : V₁ →ₗ[K] V₂} (hf : function.injective f) : + function.surjective f.dual_map := +begin + intro φ, + let f' := linear_equiv.of_injective f hf, + use subspace.dual_lift (range f) (f'.symm.dual_map φ), + ext x, + rw [linear_map.dual_map_apply, subspace.dual_lift_of_mem (mem_range_self f x), + linear_equiv.dual_map_apply], + congr' 1, + exact linear_equiv.symm_apply_apply f' x, +end + +lemma range_dual_map_eq_dual_annihilator_ker (f : V₁ →ₗ[K] V₂) : + f.dual_map.range = f.ker.dual_annihilator := +range_dual_map_eq_dual_annihilator_ker_of_subtype_range_surjective f $ + dual_map_surjective_of_injective (range f).injective_subtype + +/-- For vector spaces, `f.dual_map` is surjective if and only if `f` is injective -/ +@[simp] lemma dual_map_surjective_iff {f : V₁ →ₗ[K] V₂} : + function.surjective f.dual_map ↔ function.injective f := +by rw [← linear_map.range_eq_top, range_dual_map_eq_dual_annihilator_ker, + ← submodule.dual_annihilator_bot, subspace.dual_annihilator_inj, linear_map.ker_eq_bot] + +end linear_map + +namespace subspace +open submodule + +lemma dual_pairing_eq (W : subspace K V₁) : + W.dual_pairing = W.quot_annihilator_equiv.to_linear_map := +by { ext, refl } + +lemma dual_pairing_nondegenerate (W : subspace K V₁) : W.dual_pairing.nondegenerate := +begin + split, + { rw [linear_map.separating_left_iff_ker_eq_bot, dual_pairing_eq], + apply linear_equiv.ker, }, + { intros x h, + rw ← forall_dual_apply_eq_zero_iff K x, + intro φ, + simpa only [submodule.dual_pairing_apply, dual_lift_of_subtype] + using h (submodule.quotient.mk (W.dual_lift φ)), } +end + +lemma dual_copairing_nondegenerate (W : subspace K V₁) : W.dual_copairing.nondegenerate := +begin + split, + { rw [linear_map.separating_left_iff_ker_eq_bot, dual_copairing_eq], + apply linear_equiv.ker, }, + { rintro ⟨x⟩, + simp only [quotient.quot_mk_eq_mk, dual_copairing_apply, quotient.mk_eq_zero], + rw [← forall_mem_dual_annihilator_apply_eq_zero_iff, set_like.forall], + exact id, } +end + +-- Argument from https://math.stackexchange.com/a/2423263/172988 +lemma dual_annihilator_inf_eq (W W' : subspace K V₁) : + (W ⊓ W').dual_annihilator = W.dual_annihilator ⊔ W'.dual_annihilator := +begin + refine le_antisymm _ (sup_dual_annihilator_le_inf W W'), + let F : V₁ →ₗ[K] (V₁ ⧸ W) × (V₁ ⧸ W') := (submodule.mkq W).prod (submodule.mkq W'), + have : F.ker = W ⊓ W' := by simp only [linear_map.ker_prod, ker_mkq], + rw [← this, ← linear_map.range_dual_map_eq_dual_annihilator_ker], + intro φ, + rw [linear_map.mem_range], + rintro ⟨x, rfl⟩, + rw [submodule.mem_sup], + obtain ⟨⟨a, b⟩, rfl⟩ := (dual_prod_dual_equiv_dual K (V₁ ⧸ W) (V₁ ⧸ W')).surjective x, + obtain ⟨a', rfl⟩ := (dual_quot_equiv_dual_annihilator W).symm.surjective a, + obtain ⟨b', rfl⟩ := (dual_quot_equiv_dual_annihilator W').symm.surjective b, + use [a', a'.property, b', b'.property], + refl, +end + +-- This is also true if `V₁` is finite dimensional since one can restrict `ι` to some subtype +-- for which the infi and supr are the same. +-- +-- The obstruction to the `dual_annihilator_inf_eq` argument carrying through is that we need +-- for `module.dual R (Π (i : ι), V ⧸ W i) ≃ₗ[K] Π (i : ι), module.dual R (V ⧸ W i)`, which is not +-- true for infinite `ι`. One would need to add additional hypothesis on `W` (for example, it might +-- be true when the family is inf-closed). +lemma dual_annihilator_infi_eq {ι : Type*} [_root_.finite ι] (W : ι → subspace K V₁) : + (⨅ (i : ι), W i).dual_annihilator = (⨆ (i : ι), (W i).dual_annihilator) := +begin + unfreezingI { revert ι }, + refine finite.induction_empty_option _ _ _, + { intros α β h hyp W, + rw [← h.infi_comp, hyp (W ∘ h), ← h.supr_comp], }, + { intro W, + rw [supr_of_empty', infi_of_empty', Inf_empty, Sup_empty, dual_annihilator_top], }, + { introsI α _ h W, + rw [infi_option, supr_option, dual_annihilator_inf_eq, h], } +end + +/-- For vector spaces, dual annihilators carry direct sum decompositions +to direct sum decompositions. -/ +lemma is_compl_dual_annihilator {W W' : subspace K V₁} (h : is_compl W W') : + is_compl W.dual_annihilator W'.dual_annihilator := +begin + rw [is_compl_iff, disjoint_iff, codisjoint_iff] at h ⊢, + rw [← dual_annihilator_inf_eq, ← dual_annihilator_sup_eq, h.1, h.2, + dual_annihilator_top, dual_annihilator_bot], + exact ⟨rfl, rfl⟩ +end + +/-- For finite-dimensional vector spaces, one can distribute duals over quotients by identifying +`W.dual_lift.range` with `W`. Note that this depends on a choice of splitting of `V₁`. -/ +def dual_quot_distrib [finite_dimensional K V₁] (W : subspace K V₁) : + module.dual K (V₁ ⧸ W) ≃ₗ[K] (module.dual K V₁ ⧸ W.dual_lift.range) := +W.dual_quot_equiv_dual_annihilator.trans W.quot_dual_equiv_annihilator.symm + +end subspace + +section finite_dimensional + +open finite_dimensional linear_map variable [finite_dimensional K V₂] +namespace linear_map + +-- TODO(kmill) remove finite_dimensional if possible +-- see https://github.com/leanprover-community/mathlib/pull/17521#discussion_r1083242551 @[simp] lemma finrank_range_dual_map_eq_finrank_range (f : V₁ →ₗ[K] V₂) : finrank K f.dual_map.range = finrank K f.range := begin @@ -885,17 +1251,6 @@ begin rw [finrank_range_add_finrank_ker f.dual_map, add_comm, this], end -lemma range_dual_map_eq_dual_annihilator_ker [finite_dimensional K V₁] (f : V₁ →ₗ[K] V₂) : - f.dual_map.range = f.ker.dual_annihilator := -begin - refine eq_of_le_of_finrank_eq f.range_dual_map_le_dual_annihilator_ker _, - have := submodule.finrank_quotient_add_finrank f.ker, - rw (subspace.quot_equiv_annihilator f.ker).finrank_eq at this, - refine add_left_injective (finrank K f.ker) _, - simp_rw [this, finrank_range_dual_map_eq_finrank_range], - exact finrank_range_add_finrank_ker f, -end - /-- `f.dual_map` is injective if and only if `f` is surjective -/ @[simp] lemma dual_map_injective_iff {f : V₁ →ₗ[K] V₂} : function.injective f.dual_map ↔ function.surjective f := @@ -910,52 +1265,16 @@ begin linear_map.finrank_range_add_finrank_ker, subspace.dual_finrank_eq], end -/-- `f.dual_map` is surjective if and only if `f` is injective -/ -@[simp] lemma dual_map_surjective_iff [finite_dimensional K V₁] {f : V₁ →ₗ[K] V₂} : - function.surjective f.dual_map ↔ function.injective f := -begin - rw [← range_eq_top, ← ker_eq_bot], - split, - { intro h, - rw [← finrank_eq_zero, ← add_right_inj (finite_dimensional.finrank K f.range), - add_zero, linear_map.finrank_range_add_finrank_ker, - ← linear_map.finrank_range_dual_map_eq_finrank_range, h, - finrank_top, subspace.dual_finrank_eq], }, - { intro h, - rw [range_dual_map_eq_dual_annihilator_ker, h], - exact submodule.dual_annihilator_bot, }, -end - /-- `f.dual_map` is bijective if and only if `f` is -/ -@[simp] lemma dual_map_bijective_iff [finite_dimensional K V₁] {f : V₁ →ₗ[K] V₂} : +@[simp] lemma dual_map_bijective_iff {f : V₁ →ₗ[K] V₂} : function.bijective f.dual_map ↔ function.bijective f := by simp_rw [function.bijective, dual_map_surjective_iff, dual_map_injective_iff, and.comm] -end finite_dimensional - -section field - -variables {K V : Type*} -variables [field K] [add_comm_group V] [module K V] - -lemma dual_pairing_nondegenerate : (dual_pairing K V).nondegenerate := -begin - refine ⟨separating_left_iff_ker_eq_bot.mpr ker_id, _⟩, - intros x, - contrapose, - rintros hx : x ≠ 0, - rw [not_forall], - let f : V →ₗ[K] K := classical.some (linear_pmap.mk_span_singleton x 1 hx).to_fun.exists_extend, - use [f], - refine ne_zero_of_eq_one _, - have h : f.comp (K ∙ x).subtype = (linear_pmap.mk_span_singleton x 1 hx).to_fun := - classical.some_spec (linear_pmap.mk_span_singleton x (1 : K) hx).to_fun.exists_extend, - exact (fun_like.congr_fun h _).trans (linear_pmap.mk_span_singleton_apply _ hx _), -end +end linear_map -end field +end finite_dimensional -end linear_map +end vector_space namespace tensor_product From 468b141b14016d54b479eb7a0fff1e360b7e3cf6 Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Tue, 24 Jan 2023 03:14:41 +0000 Subject: [PATCH 0336/1291] refactor(analysis/complex/basic, data/complex/is_R_or_C): downgrade imports (#18217) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The files `data/complex/is_R_or_C` and `analysis/complex/basic` are imported widely across the analysis library: for example - `data/complex/is_R_or_C` into inner product spaces - `analysis/complex/basic` into the construction of `rpow` and thence into Lp spaces and the Bochner integral So it is useful (for the port and for compilation time) to make them as elementary as possible. Currently they both import `analysis/normed_space/operator_norm` and `analysis/normed_space/finite_dimension`, rather heavy imports, but use quite little from these files: they provide lemmas about the operator norms of `re`/`im`/`conj`/`of_real`, and get cheaply some facts about the topology of `ℂ`/`is_R_or_C` via real-finite-dimensionality. This PR splits both files. - `analysis/complex/basic` is split in place, with substantially downgraded imports, and with a few heavier lemmas moved to `analysis/complex/operator_norm`. (And a few lemmas moved earlier to `data/complex/module`.). I wrote direct proofs for the completeness and properness of `ℂ` so that these facts can remain available by importing this file, even though with heavier imports these facts can be deduced from the real-finite-dimensionality of `ℂ`. - `data/complex/is_R_or_C` is split into `data/is_R_or_C/basic` (lightweight file containing most of the former file) and `data/is_R_or_C/lemmas` (a few heavier lemmas). Also rename `equiv_real_prod_add_hom_lm` to `equiv_real_prod_lm` (an apparent typo), and `equiv_real_prodₗ` to `equiv_real_prod_clm` (also an apparent error since in our naming convention `ₗ` usually refers to linearity, not continuous-linearity). --- src/algebra/star/chsh.lean | 4 +- src/analysis/calculus/cont_diff.lean | 1 + src/analysis/calculus/mean_value.lean | 3 +- src/analysis/complex/basic.lean | 94 ++++++++----------- src/analysis/complex/cauchy_integral.lean | 2 +- src/analysis/complex/conformal.lean | 1 + src/analysis/complex/operator_norm.lean | 53 +++++++++++ src/analysis/complex/re_im_topology.lean | 14 +-- src/analysis/convex/gauge.lean | 2 +- .../inner_product_space/projection.lean | 1 + .../locally_convex/continuous_of_bounded.lean | 2 +- src/analysis/normed/group/pointwise.lean | 13 ++- src/analysis/normed_space/extend.lean | 3 +- src/analysis/normed_space/is_R_or_C.lean | 2 +- src/analysis/normed_space/star/matrix.lean | 2 +- .../special_functions/compare_exp.lean | 1 + .../special_functions/polar_coord.lean | 2 +- src/analysis/special_functions/pow.lean | 3 +- src/data/complex/module.lean | 10 ++ .../is_R_or_C.lean => is_R_or_C/basic.lean} | 70 +------------- src/data/is_R_or_C/lemmas.lean | 82 ++++++++++++++++ .../function/special_functions/basic.lean | 1 + .../measure/complex_lebesgue.lean | 2 +- src/number_theory/l_series.lean | 1 + src/number_theory/modular.lean | 1 + src/topology/continuous_function/ideals.lean | 2 +- .../stone_weierstrass.lean | 2 +- src/topology/metric_space/antilipschitz.lean | 7 ++ src/topology/metric_space/basic.lean | 10 ++ 29 files changed, 245 insertions(+), 146 deletions(-) create mode 100644 src/analysis/complex/operator_norm.lean rename src/data/{complex/is_R_or_C.lean => is_R_or_C/basic.lean} (93%) create mode 100644 src/data/is_R_or_C/lemmas.lean diff --git a/src/algebra/star/chsh.lean b/src/algebra/star/chsh.lean index 56fbc3542f280..56a574062a67e 100644 --- a/src/algebra/star/chsh.lean +++ b/src/algebra/star/chsh.lean @@ -3,8 +3,8 @@ Copyright (c) 2020 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ -import algebra.star.basic -import data.complex.is_R_or_C +import algebra.char_p.invertible +import data.real.sqrt /-! # The Clauser-Horne-Shimony-Holt inequality and Tsirelson's inequality. diff --git a/src/analysis/calculus/cont_diff.lean b/src/analysis/calculus/cont_diff.lean index 08debecca04e9..287b83ddb8304 100644 --- a/src/analysis/calculus/cont_diff.lean +++ b/src/analysis/calculus/cont_diff.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ import analysis.calculus.mean_value +import analysis.normed_space.finite_dimension import analysis.normed_space.multilinear import analysis.calculus.formal_multilinear_series import data.enat.basic diff --git a/src/analysis/calculus/mean_value.lean b/src/analysis/calculus/mean_value.lean index 636a38d50cb4e..53a1238a6040b 100644 --- a/src/analysis/calculus/mean_value.lean +++ b/src/analysis/calculus/mean_value.lean @@ -6,7 +6,8 @@ Authors: Sébastien Gouëzel, Yury Kudryashov import analysis.calculus.local_extr import analysis.convex.slope import analysis.convex.normed -import data.complex.is_R_or_C +import data.is_R_or_C.basic +import topology.instances.real_vector_space /-! # The mean value inequality and equalities diff --git a/src/analysis/complex/basic.lean b/src/analysis/complex/basic.lean index 2282d4c0ef867..e33c4d037cd5f 100644 --- a/src/analysis/complex/basic.lean +++ b/src/analysis/complex/basic.lean @@ -3,8 +3,10 @@ Copyright (c) Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ -import data.complex.determinant -import data.complex.is_R_or_C +import data.complex.module +import data.is_R_or_C.basic +import topology.algebra.module.infinite_sum +import topology.instances.real_vector_space /-! # Normed space structure on `ℂ`. @@ -28,6 +30,9 @@ isometries in `of_real_li` and `conj_lie`. We also register the fact that `ℂ` is an `is_R_or_C` field. -/ + +assert_not_exists absorbs + noncomputable theory namespace complex @@ -151,6 +156,36 @@ lemma norm_eq_one_of_pow_eq_one {ζ : ℂ} {n : ℕ} (h : ζ ^ n = 1) (hn : n ‖ζ‖ = 1 := congr_arg coe (nnnorm_eq_one_of_pow_eq_one h hn) +lemma equiv_real_prod_apply_le (z : ℂ) : ‖equiv_real_prod z‖ ≤ abs z := +by simp [prod.norm_def, abs_re_le_abs, abs_im_le_abs] + +lemma equiv_real_prod_apply_le' (z : ℂ) : ‖equiv_real_prod z‖ ≤ 1 * abs z := +by simpa using equiv_real_prod_apply_le z + +lemma lipschitz_equiv_real_prod : lipschitz_with 1 equiv_real_prod := +by simpa using + add_monoid_hom_class.lipschitz_of_bound equiv_real_prod_lm 1 equiv_real_prod_apply_le' + +lemma antilipschitz_equiv_real_prod : antilipschitz_with (nnreal.sqrt 2) equiv_real_prod := +by simpa using + add_monoid_hom_class.antilipschitz_of_bound equiv_real_prod_lm abs_le_sqrt_two_mul_max + +lemma uniform_embedding_equiv_real_prod : uniform_embedding equiv_real_prod := +antilipschitz_equiv_real_prod.uniform_embedding lipschitz_equiv_real_prod.uniform_continuous + +instance : complete_space ℂ := +(complete_space_congr uniform_embedding_equiv_real_prod).mpr infer_instance + +/-- The natural `continuous_linear_equiv` from `ℂ` to `ℝ × ℝ`. -/ +@[simps apply symm_apply_re symm_apply_im { simp_rhs := tt }] +def equiv_real_prod_clm : ℂ ≃L[ℝ] ℝ × ℝ := +equiv_real_prod_lm.to_continuous_linear_equiv_of_bounds 1 (real.sqrt 2) +equiv_real_prod_apply_le' +(λ p, abs_le_sqrt_two_mul_max (equiv_real_prod.symm p)) + +instance : proper_space ℂ := +(id lipschitz_equiv_real_prod : lipschitz_with 1 equiv_real_prod_clm.to_homeomorph).proper_space + /-- The `abs` function on `ℂ` is proper. -/ lemma tendsto_abs_cocompact_at_top : filter.tendsto abs (filter.cocompact ℂ) filter.at_top := tendsto_norm_cocompact_at_top @@ -172,13 +207,6 @@ def re_clm : ℂ →L[ℝ] ℝ := re_lm.mk_continuous 1 (λ x, by simp [abs_re_l @[simp] lemma re_clm_apply (z : ℂ) : (re_clm : ℂ → ℝ) z = z.re := rfl -@[simp] lemma re_clm_norm : ‖re_clm‖ = 1 := -le_antisymm (linear_map.mk_continuous_norm_le _ zero_le_one _) $ -calc 1 = ‖re_clm 1‖ : by simp - ... ≤ ‖re_clm‖ : unit_le_op_norm _ _ (by simp) - -@[simp] lemma re_clm_nnnorm : ‖re_clm‖₊ = 1 := subtype.ext re_clm_norm - /-- Continuous linear map version of the real part function, from `ℂ` to `ℝ`. -/ def im_clm : ℂ →L[ℝ] ℝ := im_lm.mk_continuous 1 (λ x, by simp [abs_im_le_abs]) @@ -188,13 +216,6 @@ def im_clm : ℂ →L[ℝ] ℝ := im_lm.mk_continuous 1 (λ x, by simp [abs_im_l @[simp] lemma im_clm_apply (z : ℂ) : (im_clm : ℂ → ℝ) z = z.im := rfl -@[simp] lemma im_clm_norm : ‖im_clm‖ = 1 := -le_antisymm (linear_map.mk_continuous_norm_le _ zero_le_one _) $ -calc 1 = ‖im_clm I‖ : by simp - ... ≤ ‖im_clm‖ : unit_le_op_norm _ _ (by simp) - -@[simp] lemma im_clm_nnnorm : ‖im_clm‖₊ = 1 := subtype.ext im_clm_norm - lemma restrict_scalars_one_smul_right' (x : E) : continuous_linear_map.restrict_scalars ℝ ((1 : ℂ →L[ℂ] ℂ).smul_right x : ℂ →L[ℂ] E) = re_clm.smul_right x + I • im_clm.smul_right x := @@ -225,14 +246,6 @@ by rw [← dist_conj_conj, conj_conj] lemma nndist_conj_comm (z w : ℂ) : nndist (conj z) w = nndist z (conj w) := subtype.ext $ dist_conj_comm _ _ -/-- The determinant of `conj_lie`, as a linear map. -/ -@[simp] lemma det_conj_lie : (conj_lie.to_linear_equiv : ℂ →ₗ[ℝ] ℂ).det = -1 := -det_conj_ae - -/-- The determinant of `conj_lie`, as a linear equiv. -/ -@[simp] lemma linear_equiv_det_conj_lie : conj_lie.to_linear_equiv.det = -1 := -linear_equiv_det_conj_ae - instance : has_continuous_star ℂ := ⟨conj_lie.continuous⟩ @[continuity] lemma continuous_conj : continuous (conj : ℂ → ℂ) := continuous_star @@ -253,11 +266,6 @@ def conj_cle : ℂ ≃L[ℝ] ℂ := conj_lie @[simp] lemma conj_cle_apply (z : ℂ) : conj_cle z = conj z := rfl -@[simp] lemma conj_cle_norm : ‖(conj_cle : ℂ →L[ℝ] ℂ)‖ = 1 := -conj_lie.to_linear_isometry.norm_to_continuous_linear_map - -@[simp] lemma conj_cle_nnorm : ‖(conj_cle : ℂ →L[ℝ] ℂ)‖₊ = 1 := subtype.ext conj_cle_norm - /-- Linear isometry version of the canonical embedding of `ℝ` in `ℂ`. -/ def of_real_li : ℝ →ₗᵢ[ℝ] ℂ := ⟨of_real_am.to_linear_map, norm_real⟩ @@ -281,10 +289,6 @@ def of_real_clm : ℝ →L[ℝ] ℂ := of_real_li.to_continuous_linear_map @[simp] lemma of_real_clm_apply (x : ℝ) : of_real_clm x = x := rfl -@[simp] lemma of_real_clm_norm : ‖of_real_clm‖ = 1 := of_real_li.norm_to_continuous_linear_map - -@[simp] lemma of_real_clm_nnnorm : ‖of_real_clm‖₊ = 1 := subtype.ext $ of_real_clm_norm - noncomputable instance : is_R_or_C ℂ := { re := ⟨complex.re, complex.zero_re, complex.add_re⟩, im := ⟨complex.im, complex.zero_im, complex.add_im⟩, @@ -320,28 +324,6 @@ by rw [eq_re_of_real_le hz, is_R_or_C.norm_of_real, real.norm_of_nonneg (complex end complex_order -section - -variables {α β γ : Type*} - [add_comm_monoid α] [topological_space α] [add_comm_monoid γ] [topological_space γ] - -/-- The natural `add_equiv` from `ℂ` to `ℝ × ℝ`. -/ -@[simps apply symm_apply_re symm_apply_im { simp_rhs := tt }] -def equiv_real_prod_add_hom : ℂ ≃+ ℝ × ℝ := -{ map_add' := by simp, .. equiv_real_prod } - -/-- The natural `linear_equiv` from `ℂ` to `ℝ × ℝ`. -/ -@[simps apply symm_apply_re symm_apply_im { simp_rhs := tt }] -def equiv_real_prod_add_hom_lm : ℂ ≃ₗ[ℝ] ℝ × ℝ := -{ map_smul' := by simp [equiv_real_prod_add_hom], .. equiv_real_prod_add_hom } - -/-- The natural `continuous_linear_equiv` from `ℂ` to `ℝ × ℝ`. -/ -@[simps apply symm_apply_re symm_apply_im { simp_rhs := tt }] -def equiv_real_prodₗ : ℂ ≃L[ℝ] ℝ × ℝ := -equiv_real_prod_add_hom_lm.to_continuous_linear_equiv - -end - lemma has_sum_iff {α} (f : α → ℂ) (c : ℂ) : has_sum f c ↔ has_sum (λ x, (f x).re) c.re ∧ has_sum (λ x, (f x).im) c.im := begin @@ -349,7 +331,7 @@ begin -- `has_sum.mapL` here: refine ⟨λ h, ⟨re_clm.has_sum h, im_clm.has_sum h⟩, _⟩, rintro ⟨h₁, h₂⟩, - convert (h₁.prod_mk h₂).mapL equiv_real_prodₗ.symm.to_continuous_linear_map, + convert (h₁.prod_mk h₂).mapL equiv_real_prod_clm.symm.to_continuous_linear_map, { ext x; refl }, { cases c, refl } end diff --git a/src/analysis/complex/cauchy_integral.lean b/src/analysis/complex/cauchy_integral.lean index b4c6d6e96f5ee..6a123025b7ec6 100644 --- a/src/analysis/complex/cauchy_integral.lean +++ b/src/analysis/complex/cauchy_integral.lean @@ -169,7 +169,7 @@ lemma integral_boundary_rect_of_has_fderiv_at_real_off_countable (f : ℂ → E) (I • ∫ y : ℝ in z.im..w.im, f (re w + y * I)) - I • ∫ y : ℝ in z.im..w.im, f (re z + y * I) = ∫ x : ℝ in z.re..w.re, ∫ y : ℝ in z.im..w.im, I • f' (x + y * I) 1 - f' (x + y * I) I := begin - set e : (ℝ × ℝ) ≃L[ℝ] ℂ := equiv_real_prodₗ.symm, + set e : (ℝ × ℝ) ≃L[ℝ] ℂ := equiv_real_prod_clm.symm, have he : ∀ x y : ℝ, ↑x + ↑y * I = e (x, y), from λ x y, (mk_eq_add_mul_I x y).symm, have he₁ : e (1, 0) = 1 := rfl, have he₂ : e (0, 1) = I := rfl, simp only [he] at *, diff --git a/src/analysis/complex/conformal.lean b/src/analysis/complex/conformal.lean index 33c5a834d814a..895a1c2cc86c3 100644 --- a/src/analysis/complex/conformal.lean +++ b/src/analysis/complex/conformal.lean @@ -5,6 +5,7 @@ Authors: Yourong Zang -/ import analysis.complex.isometry import analysis.normed_space.conformal_linear_map +import analysis.normed_space.finite_dimension /-! # Conformal maps between complex vector spaces diff --git a/src/analysis/complex/operator_norm.lean b/src/analysis/complex/operator_norm.lean new file mode 100644 index 0000000000000..a0133388f963b --- /dev/null +++ b/src/analysis/complex/operator_norm.lean @@ -0,0 +1,53 @@ +/- +Copyright (c) Sébastien Gouëzel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sébastien Gouëzel +-/ +import analysis.complex.basic +import analysis.normed_space.operator_norm +import data.complex.determinant + +/-! # The basic continuous linear maps associated to `ℂ` + +The continuous linear maps `complex.re_clm` (real part), `complex.im_clm` (imaginary part), +`complex.conj_cle` (conjugation), and `complex.of_real_clm` (inclusion of `ℝ`) were introduced in +`analysis.complex.operator_norm`. This file contains a few calculations requiring more imports: +the operator norm and (for `complex.conj_cle`) the determinant. +-/ + +open continuous_linear_map + +namespace complex + +/-- The determinant of `conj_lie`, as a linear map. -/ +@[simp] lemma det_conj_lie : (conj_lie.to_linear_equiv : ℂ →ₗ[ℝ] ℂ).det = -1 := +det_conj_ae + +/-- The determinant of `conj_lie`, as a linear equiv. -/ +@[simp] lemma linear_equiv_det_conj_lie : conj_lie.to_linear_equiv.det = -1 := +linear_equiv_det_conj_ae + +@[simp] lemma re_clm_norm : ‖re_clm‖ = 1 := +le_antisymm (linear_map.mk_continuous_norm_le _ zero_le_one _) $ +calc 1 = ‖re_clm 1‖ : by simp + ... ≤ ‖re_clm‖ : unit_le_op_norm _ _ (by simp) + +@[simp] lemma re_clm_nnnorm : ‖re_clm‖₊ = 1 := subtype.ext re_clm_norm + +@[simp] lemma im_clm_norm : ‖im_clm‖ = 1 := +le_antisymm (linear_map.mk_continuous_norm_le _ zero_le_one _) $ +calc 1 = ‖im_clm I‖ : by simp + ... ≤ ‖im_clm‖ : unit_le_op_norm _ _ (by simp) + +@[simp] lemma im_clm_nnnorm : ‖im_clm‖₊ = 1 := subtype.ext im_clm_norm + +@[simp] lemma conj_cle_norm : ‖(conj_cle : ℂ →L[ℝ] ℂ)‖ = 1 := +conj_lie.to_linear_isometry.norm_to_continuous_linear_map + +@[simp] lemma conj_cle_nnorm : ‖(conj_cle : ℂ →L[ℝ] ℂ)‖₊ = 1 := subtype.ext conj_cle_norm + +@[simp] lemma of_real_clm_norm : ‖of_real_clm‖ = 1 := of_real_li.norm_to_continuous_linear_map + +@[simp] lemma of_real_clm_nnnorm : ‖of_real_clm‖₊ = 1 := subtype.ext $ of_real_clm_norm + +end complex diff --git a/src/analysis/complex/re_im_topology.lean b/src/analysis/complex/re_im_topology.lean index 6eec8014fa1c9..5410395907c29 100644 --- a/src/analysis/complex/re_im_topology.lean +++ b/src/analysis/complex/re_im_topology.lean @@ -38,11 +38,11 @@ namespace complex /-- `complex.re` turns `ℂ` into a trivial topological fiber bundle over `ℝ`. -/ lemma is_homeomorphic_trivial_fiber_bundle_re : is_homeomorphic_trivial_fiber_bundle ℝ re := -⟨equiv_real_prodₗ.to_homeomorph, λ z, rfl⟩ +⟨equiv_real_prod_clm.to_homeomorph, λ z, rfl⟩ /-- `complex.im` turns `ℂ` into a trivial topological fiber bundle over `ℝ`. -/ lemma is_homeomorphic_trivial_fiber_bundle_im : is_homeomorphic_trivial_fiber_bundle ℝ im := -⟨equiv_real_prodₗ.to_homeomorph.trans (homeomorph.prod_comm ℝ ℝ), λ z, rfl⟩ +⟨equiv_real_prod_clm.to_homeomorph.trans (homeomorph.prod_comm ℝ ℝ), λ z, rfl⟩ lemma is_open_map_re : is_open_map re := is_homeomorphic_trivial_fiber_bundle_re.is_open_map_proj lemma is_open_map_im : is_open_map im := is_homeomorphic_trivial_fiber_bundle_im.is_open_map_proj @@ -117,8 +117,8 @@ by simpa only [frontier_Ioi] using frontier_preimage_re (Ioi a) by simpa only [frontier_Ioi] using frontier_preimage_im (Ioi a) lemma closure_re_prod_im (s t : set ℝ) : closure (s ×ℂ t) = closure s ×ℂ closure t := -by simpa only [← preimage_eq_preimage equiv_real_prodₗ.symm.to_homeomorph.surjective, - equiv_real_prodₗ.symm.to_homeomorph.preimage_closure] +by simpa only [← preimage_eq_preimage equiv_real_prod_clm.symm.to_homeomorph.surjective, + equiv_real_prod_clm.symm.to_homeomorph.preimage_closure] using @closure_prod_eq _ _ _ _ s t lemma interior_re_prod_im (s t : set ℝ) : interior (s ×ℂ t) = interior s ×ℂ interior t := @@ -126,8 +126,8 @@ by rw [re_prod_im, re_prod_im, interior_inter, interior_preimage_re, interior_pr lemma frontier_re_prod_im (s t : set ℝ) : frontier (s ×ℂ t) = (closure s ×ℂ frontier t) ∪ (frontier s ×ℂ closure t) := -by simpa only [← preimage_eq_preimage equiv_real_prodₗ.symm.to_homeomorph.surjective, - equiv_real_prodₗ.symm.to_homeomorph.preimage_frontier] +by simpa only [← preimage_eq_preimage equiv_real_prod_clm.symm.to_homeomorph.surjective, + equiv_real_prod_clm.symm.to_homeomorph.preimage_frontier] using frontier_prod_eq s t lemma frontier_set_of_le_re_and_le_im (a b : ℝ) : @@ -152,4 +152,4 @@ lemma is_closed.re_prod_im (hs : is_closed s) (ht : is_closed t) : is_closed (s (hs.preimage continuous_re).inter (ht.preimage continuous_im) lemma metric.bounded.re_prod_im (hs : bounded s) (ht : bounded t) : bounded (s ×ℂ t) := -equiv_real_prodₗ.antilipschitz.bounded_preimage (hs.prod ht) +antilipschitz_equiv_real_prod.bounded_preimage (hs.prod ht) diff --git a/src/analysis/convex/gauge.lean b/src/analysis/convex/gauge.lean index 2f53e33cfa4c7..62cc0de3f37c7 100644 --- a/src/analysis/convex/gauge.lean +++ b/src/analysis/convex/gauge.lean @@ -6,7 +6,7 @@ Authors: Yaël Dillies, Bhavik Mehta import analysis.convex.basic import analysis.normed_space.pointwise import analysis.seminorm -import data.complex.is_R_or_C +import data.is_R_or_C.basic import tactic.congrm /-! diff --git a/src/analysis/inner_product_space/projection.lean b/src/analysis/inner_product_space/projection.lean index 482ec4b838a5b..aa7d538295651 100644 --- a/src/analysis/inner_product_space/projection.lean +++ b/src/analysis/inner_product_space/projection.lean @@ -6,6 +6,7 @@ Authors: Zhouhang Zhou, Frédéric Dupuis, Heather Macbeth import analysis.convex.basic import analysis.inner_product_space.symmetric import analysis.normed_space.is_R_or_C +import data.is_R_or_C.lemmas /-! # The orthogonal projection diff --git a/src/analysis/locally_convex/continuous_of_bounded.lean b/src/analysis/locally_convex/continuous_of_bounded.lean index 0bfc937e9b80f..773fb5b28b056 100644 --- a/src/analysis/locally_convex/continuous_of_bounded.lean +++ b/src/analysis/locally_convex/continuous_of_bounded.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Moritz Doll -/ import analysis.locally_convex.bounded -import data.complex.is_R_or_C +import data.is_R_or_C.basic /-! # Continuity and Von Neumann boundedness diff --git a/src/analysis/normed/group/pointwise.lean b/src/analysis/normed/group/pointwise.lean index 8cb7806fbb3da..8dde43dd8e408 100644 --- a/src/analysis/normed/group/pointwise.lean +++ b/src/analysis/normed/group/pointwise.lean @@ -122,9 +122,20 @@ lemma closed_ball_one_mul_singleton : closed_ball 1 δ * {x} = closed_ball x δ @[to_additive] lemma closed_ball_one_div_singleton : closed_ball 1 δ / {x} = closed_ball x⁻¹ δ := by simp -@[simp, to_additive] lemma smul_closed_ball_one : x • closed_ball 1 δ = closed_ball x δ := +-- This is the `to_additive` version of the below, but it will later follow as a special case of +-- `vadd_closed_ball` for `normed_add_torsor`s, so we give it higher simp priority. +-- (There is no `normed_mul_torsor`, hence the asymmetry between additive and multiplicative +-- versions.) +@[simp, priority 1100] lemma vadd_closed_ball_zero {E : Type*} [seminormed_add_comm_group E] (δ : ℝ) + (x : E) : + x +ᵥ metric.closed_ball 0 δ = metric.closed_ball x δ := +by { ext, simp [mem_vadd_set_iff_neg_vadd_mem, neg_add_eq_sub, dist_eq_norm_sub] } + +@[simp] lemma smul_closed_ball_one : x • closed_ball 1 δ = closed_ball x δ := by { ext, simp [mem_smul_set_iff_inv_smul_mem, inv_mul_eq_div, dist_eq_norm_div] } +attribute [to_additive] smul_closed_ball_one + @[to_additive] lemma mul_ball_one : s * ball 1 δ = thickening δ s := begin rw thickening_eq_bUnion_ball, diff --git a/src/analysis/normed_space/extend.lean b/src/analysis/normed_space/extend.lean index 90ea75b806679..f0a198fc8ed45 100644 --- a/src/analysis/normed_space/extend.lean +++ b/src/analysis/normed_space/extend.lean @@ -4,8 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Ruben Van de Velde -/ +import analysis.normed_space.operator_norm import algebra.algebra.restrict_scalars -import data.complex.is_R_or_C +import data.is_R_or_C.basic /-! # Extending a continuous `ℝ`-linear map to a continuous `𝕜`-linear map diff --git a/src/analysis/normed_space/is_R_or_C.lean b/src/analysis/normed_space/is_R_or_C.lean index 149808aefc381..b2941c8b4cc54 100644 --- a/src/analysis/normed_space/is_R_or_C.lean +++ b/src/analysis/normed_space/is_R_or_C.lean @@ -3,7 +3,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ä -/ -import data.complex.is_R_or_C +import data.is_R_or_C.basic import analysis.normed_space.operator_norm import analysis.normed_space.pointwise diff --git a/src/analysis/normed_space/star/matrix.lean b/src/analysis/normed_space/star/matrix.lean index 241239a3b3976..035eb8318992f 100644 --- a/src/analysis/normed_space/star/matrix.lean +++ b/src/analysis/normed_space/star/matrix.lean @@ -5,7 +5,7 @@ Authors: Hans Parshall -/ import analysis.matrix import analysis.normed_space.basic -import data.complex.is_R_or_C +import data.is_R_or_C.basic import linear_algebra.unitary_group /-! diff --git a/src/analysis/special_functions/compare_exp.lean b/src/analysis/special_functions/compare_exp.lean index 653f01558aa33..711f1963e9170 100644 --- a/src/analysis/special_functions/compare_exp.lean +++ b/src/analysis/special_functions/compare_exp.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import analysis.special_functions.pow +import analysis.asymptotics.asymptotic_equivalent import analysis.asymptotics.specific_asymptotics /-! diff --git a/src/analysis/special_functions/polar_coord.lean b/src/analysis/special_functions/polar_coord.lean index 8b76bf9a13ab3..81c941f14e955 100644 --- a/src/analysis/special_functions/polar_coord.lean +++ b/src/analysis/special_functions/polar_coord.lean @@ -91,7 +91,7 @@ It is a homeomorphism between `ℝ^2 - (-∞, 0]` and `(0, +∞) × (-π, π)`. { rintros ⟨x, y⟩ hxy, simpa only using hxy }, apply continuous_on.comp (λ z hz, _) _ A, { exact (complex.continuous_at_arg hz).continuous_within_at }, - { exact complex.equiv_real_prodₗ.symm.continuous.continuous_on } + { exact complex.equiv_real_prod_clm.symm.continuous.continuous_on } end } lemma has_fderiv_at_polar_coord_symm (p : ℝ × ℝ) : diff --git a/src/analysis/special_functions/pow.lean b/src/analysis/special_functions/pow.lean index ddcfb486c74dd..aa46f782dbe80 100644 --- a/src/analysis/special_functions/pow.lean +++ b/src/analysis/special_functions/pow.lean @@ -20,8 +20,7 @@ We also prove basic properties of these functions. noncomputable theory -open_locale classical real topological_space nnreal ennreal filter big_operators asymptotics -open_locale complex_conjugate +open_locale classical real topological_space nnreal ennreal filter big_operators complex_conjugate open filter finset set namespace complex diff --git a/src/data/complex/module.lean b/src/data/complex/module.lean index a79d585495917..6da5bc73d98c2 100644 --- a/src/data/complex/module.lean +++ b/src/data/complex/module.lean @@ -266,6 +266,16 @@ begin exacts [h, conj_I.symm ▸ h], end +/-- The natural `add_equiv` from `ℂ` to `ℝ × ℝ`. -/ +@[simps apply symm_apply_re symm_apply_im { simp_rhs := tt }] +def equiv_real_prod_add_hom : ℂ ≃+ ℝ × ℝ := +{ map_add' := by simp, .. equiv_real_prod } + +/-- The natural `linear_equiv` from `ℂ` to `ℝ × ℝ`. -/ +@[simps apply symm_apply_re symm_apply_im { simp_rhs := tt }] +def equiv_real_prod_lm : ℂ ≃ₗ[ℝ] ℝ × ℝ := +{ map_smul' := by simp [equiv_real_prod_add_hom], .. equiv_real_prod_add_hom } + section lift variables {A : Type*} [ring A] [algebra ℝ A] diff --git a/src/data/complex/is_R_or_C.lean b/src/data/is_R_or_C/basic.lean similarity index 93% rename from src/data/complex/is_R_or_C.lean rename to src/data/is_R_or_C/basic.lean index b9a613eb44f78..d0508ec16e351 100644 --- a/src/data/complex/is_R_or_C.lean +++ b/src/data/is_R_or_C/basic.lean @@ -4,9 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Frédéric Dupuis -/ import data.real.sqrt -import field_theory.tower -import analysis.normed_space.finite_dimension import analysis.normed_space.star.basic +import analysis.normed_space.continuous_linear_map /-! # `is_R_or_C`: a typeclass for ℝ or ℂ @@ -34,6 +33,8 @@ in `data/nat/cast`. See also Note [coercion into rings] for more details. In addition, several lemmas need to be set at priority 900 to make sure that they do not override their counterparts in `complex.lean` (which causes linter errors). + +A few lemmas requiring heavier imports are in `data.is_R_or_C.lemmas`. -/ open_locale big_operators @@ -667,57 +668,6 @@ ring_hom.map_finsupp_prod _ f g end is_R_or_C -namespace polynomial - -open_locale polynomial - -lemma of_real_eval (p : ℝ[X]) (x : ℝ) : (p.eval x : K) = aeval ↑x p := -(@aeval_algebra_map_apply_eq_algebra_map_eval ℝ K _ _ _ x p).symm - -end polynomial - -namespace finite_dimensional - -open_locale classical -open is_R_or_C - -/-- This instance generates a type-class problem with a metavariable `?m` that should satisfy -`is_R_or_C ?m`. Since this can only be satisfied by `ℝ` or `ℂ`, this does not cause problems. -/ -library_note "is_R_or_C instance" - -/-- An `is_R_or_C` field is finite-dimensional over `ℝ`, since it is spanned by `{1, I}`. -/ -@[nolint dangerous_instance] instance is_R_or_C_to_real : finite_dimensional ℝ K := -⟨⟨{1, I}, - begin - rw eq_top_iff, - intros a _, - rw [finset.coe_insert, finset.coe_singleton, submodule.mem_span_insert], - refine ⟨re a, (im a) • I, _, _⟩, - { rw submodule.mem_span_singleton, - use im a }, - simp [re_add_im a, algebra.smul_def, algebra_map_eq_of_real] - end⟩⟩ - -variables (K E) [normed_add_comm_group E] [normed_space K E] - -/-- A finite dimensional vector space over an `is_R_or_C` is a proper metric space. - -This is not an instance because it would cause a search for `finite_dimensional ?x E` before -`is_R_or_C ?x`. -/ -lemma proper_is_R_or_C [finite_dimensional K E] : proper_space E := -begin - letI : normed_space ℝ E := restrict_scalars.normed_space ℝ K E, - letI : finite_dimensional ℝ E := finite_dimensional.trans ℝ K E, - apply_instance -end - -variable {E} - -instance is_R_or_C.proper_space_submodule (S : submodule K E) [finite_dimensional K ↥S] : - proper_space S := -proper_is_R_or_C K S - -end finite_dimensional section instances @@ -786,14 +736,6 @@ noncomputable def re_clm : K →L[ℝ] ℝ := linear_map.mk_continuous re_lm 1 $ by { simp only [norm_eq_abs, re_lm_coe, one_mul, abs_to_real], exact abs_re_le_abs, } -@[simp, is_R_or_C_simps] lemma re_clm_norm : ‖(re_clm : K →L[ℝ] ℝ)‖ = 1 := -begin - apply le_antisymm (linear_map.mk_continuous_norm_le _ zero_le_one _), - convert continuous_linear_map.ratio_le_op_norm _ (1 : K), - { simp }, - { apply_instance } -end - @[simp, is_R_or_C_simps, norm_cast] lemma re_clm_coe : ((re_clm : K →L[ℝ] ℝ) : K →ₗ[ℝ] ℝ) = re_lm := rfl @@ -843,9 +785,6 @@ noncomputable def conj_cle : K ≃L[ℝ] K := @conj_lie K _ @[simp, is_R_or_C_simps] lemma conj_cle_apply : (conj_cle : K → K) = conj := rfl -@[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 - @[priority 100] instance : has_continuous_star K := ⟨conj_lie.continuous⟩ @@ -870,9 +809,6 @@ noncomputable def of_real_clm : ℝ →L[ℝ] K := of_real_li.to_continuous_line @[simp, is_R_or_C_simps] lemma of_real_clm_apply : (of_real_clm : ℝ → K) = coe := rfl -@[simp, is_R_or_C_simps] lemma of_real_clm_norm : ‖(of_real_clm : ℝ →L[ℝ] K)‖ = 1 := -linear_isometry.norm_to_continuous_linear_map of_real_li - @[continuity] lemma continuous_of_real : continuous (coe : ℝ → K) := of_real_li.continuous @[continuity] lemma continuous_abs : continuous (@is_R_or_C.abs K _) := diff --git a/src/data/is_R_or_C/lemmas.lean b/src/data/is_R_or_C/lemmas.lean new file mode 100644 index 0000000000000..f4685686068ff --- /dev/null +++ b/src/data/is_R_or_C/lemmas.lean @@ -0,0 +1,82 @@ +/- +Copyright (c) 2020 Frédéric Dupuis. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Frédéric Dupuis +-/ +import analysis.normed_space.finite_dimension +import field_theory.tower +import data.is_R_or_C.basic + +/-! # Further lemmas about `is_R_or_C` -/ + +variables {K E : Type*} [is_R_or_C K] + +namespace polynomial + +open_locale polynomial + +lemma of_real_eval (p : ℝ[X]) (x : ℝ) : (p.eval x : K) = aeval ↑x p := +(@aeval_algebra_map_apply_eq_algebra_map_eval ℝ K _ _ _ x p).symm + +end polynomial + +namespace finite_dimensional + +open_locale classical +open is_R_or_C + +/-- This instance generates a type-class problem with a metavariable `?m` that should satisfy +`is_R_or_C ?m`. Since this can only be satisfied by `ℝ` or `ℂ`, this does not cause problems. -/ +library_note "is_R_or_C instance" + +/-- An `is_R_or_C` field is finite-dimensional over `ℝ`, since it is spanned by `{1, I}`. -/ +@[nolint dangerous_instance] instance is_R_or_C_to_real : finite_dimensional ℝ K := +⟨⟨{1, I}, + begin + rw eq_top_iff, + intros a _, + rw [finset.coe_insert, finset.coe_singleton, submodule.mem_span_insert], + refine ⟨re a, (im a) • I, _, _⟩, + { rw submodule.mem_span_singleton, + use im a }, + simp [re_add_im a, algebra.smul_def, algebra_map_eq_of_real] + end⟩⟩ + +variables (K E) [normed_add_comm_group E] [normed_space K E] + +/-- A finite dimensional vector space over an `is_R_or_C` is a proper metric space. + +This is not an instance because it would cause a search for `finite_dimensional ?x E` before +`is_R_or_C ?x`. -/ +lemma proper_is_R_or_C [finite_dimensional K E] : proper_space E := +begin + letI : normed_space ℝ E := restrict_scalars.normed_space ℝ K E, + letI : finite_dimensional ℝ E := finite_dimensional.trans ℝ K E, + apply_instance +end + +variable {E} + +instance is_R_or_C.proper_space_submodule (S : submodule K E) [finite_dimensional K ↥S] : + proper_space S := +proper_is_R_or_C K S + +end finite_dimensional + +namespace is_R_or_C + +@[simp, is_R_or_C_simps] lemma re_clm_norm : ‖(re_clm : K →L[ℝ] ℝ)‖ = 1 := +begin + apply le_antisymm (linear_map.mk_continuous_norm_le _ zero_le_one _), + convert continuous_linear_map.ratio_le_op_norm _ (1 : K), + { simp }, + { apply_instance } +end + +@[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 + +@[simp, is_R_or_C_simps] lemma of_real_clm_norm : ‖(of_real_clm : ℝ →L[ℝ] K)‖ = 1 := +linear_isometry.norm_to_continuous_linear_map of_real_li + +end is_R_or_C diff --git a/src/measure_theory/function/special_functions/basic.lean b/src/measure_theory/function/special_functions/basic.lean index b2d848d277d4e..e394cccf6cd95 100644 --- a/src/measure_theory/function/special_functions/basic.lean +++ b/src/measure_theory/function/special_functions/basic.lean @@ -5,6 +5,7 @@ Authors: Yury Kudryashov -/ import analysis.special_functions.pow +import data.is_R_or_C.lemmas import measure_theory.constructions.borel_space /-! diff --git a/src/measure_theory/measure/complex_lebesgue.lean b/src/measure_theory/measure/complex_lebesgue.lean index 5022bc9ea59bc..dc00eaa45138e 100644 --- a/src/measure_theory/measure/complex_lebesgue.lean +++ b/src/measure_theory/measure/complex_lebesgue.lean @@ -30,7 +30,7 @@ basis_one_I.equiv_fun.to_continuous_linear_equiv.to_homeomorph.to_measurable_equ /-- Measurable equivalence between `ℂ` and `ℝ × ℝ`. -/ def measurable_equiv_real_prod : ℂ ≃ᵐ (ℝ × ℝ) := -equiv_real_prodₗ.to_homeomorph.to_measurable_equiv +equiv_real_prod_clm.to_homeomorph.to_measurable_equiv lemma volume_preserving_equiv_pi : measure_preserving measurable_equiv_pi := diff --git a/src/number_theory/l_series.lean b/src/number_theory/l_series.lean index 47fd7d2cd9d66..c40f8bde46176 100644 --- a/src/number_theory/l_series.lean +++ b/src/number_theory/l_series.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Aaron Anderson. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Aaron Anderson -/ +import analysis.normed_space.finite_dimension import analysis.p_series import number_theory.arithmetic_function import topology.algebra.infinite_sum diff --git a/src/number_theory/modular.lean b/src/number_theory/modular.lean index a91971dd18130..bf0d3051973be 100644 --- a/src/number_theory/modular.lean +++ b/src/number_theory/modular.lean @@ -5,6 +5,7 @@ Authors: Alex Kontorovich, Heather Macbeth, Marc Masdeu -/ import analysis.complex.upper_half_plane.basic +import analysis.normed_space.finite_dimension import linear_algebra.general_linear_group /-! diff --git a/src/topology/continuous_function/ideals.lean b/src/topology/continuous_function/ideals.lean index 0ca31fed642c9..b51e302e6c6f7 100644 --- a/src/topology/continuous_function/ideals.lean +++ b/src/topology/continuous_function/ideals.lean @@ -7,7 +7,7 @@ Authors: Jireh Loreaux import topology.algebra.algebra import topology.continuous_function.compact import topology.urysohns_lemma -import data.complex.is_R_or_C +import data.is_R_or_C.basic import analysis.normed_space.units import topology.algebra.module.character_space diff --git a/src/topology/continuous_function/stone_weierstrass.lean b/src/topology/continuous_function/stone_weierstrass.lean index 581de93c3648e..9756dccbdb3c3 100644 --- a/src/topology/continuous_function/stone_weierstrass.lean +++ b/src/topology/continuous_function/stone_weierstrass.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison, Heather Macbeth -/ import topology.continuous_function.weierstrass -import data.complex.is_R_or_C +import data.is_R_or_C.basic /-! # The Stone-Weierstrass theorem diff --git a/src/topology/metric_space/antilipschitz.lean b/src/topology/metric_space/antilipschitz.lean index 05db82f6f4b9a..f07c0959775e0 100644 --- a/src/topology/metric_space/antilipschitz.lean +++ b/src/topology/metric_space/antilipschitz.lean @@ -222,3 +222,10 @@ lemma lipschitz_with.to_right_inverse [pseudo_emetric_space α] [pseudo_emetric_ {f : α → β} (hf : lipschitz_with K f) {g : β → α} (hg : function.right_inverse g f) : antilipschitz_with K g := λ x y, by simpa only [hg _] using hf (g x) (g y) + +/-- The preimage of a proper space under a Lipschitz homeomorphism is proper. -/ +@[protected] +theorem lipschitz_with.proper_space [pseudo_metric_space α] [metric_space β] [proper_space β] + {K : ℝ≥0} {f : α ≃ₜ β} (hK : lipschitz_with K f) : + proper_space α := +(hK.to_right_inverse f.right_inv).proper_space f.symm.continuous f.symm.surjective diff --git a/src/topology/metric_space/basic.lean b/src/topology/metric_space/basic.lean index c00763637c991..33b8f2bdc72d5 100644 --- a/src/topology/metric_space/basic.lean +++ b/src/topology/metric_space/basic.lean @@ -2026,6 +2026,16 @@ instance complete_of_proper [proper_space α] : complete_space α := exact ⟨y, hy⟩ end⟩ +/-- A binary product of proper spaces is proper. -/ +instance prod_proper_space {α : Type*} {β : Type*} [pseudo_metric_space α] [pseudo_metric_space β] + [proper_space α] [proper_space β] : + proper_space (α × β) := +{ is_compact_closed_ball := begin + rintros ⟨x, y⟩ r, + rw ← closed_ball_prod_same x y, + apply (is_compact_closed_ball x r).prod (is_compact_closed_ball y r), + end } + /-- A finite product of proper spaces is proper. -/ instance pi_proper_space {π : β → Type*} [fintype β] [∀b, pseudo_metric_space (π b)] [h : ∀b, proper_space (π b)] : proper_space (Πb, π b) := From 0be37584a9ebeb9ee8afbe8ea52d8a4edd95d7ac Mon Sep 17 00:00:00 2001 From: leanprover-community-bot Date: Tue, 24 Jan 2023 05:14:58 +0000 Subject: [PATCH 0337/1291] chore(scripts): update nolints.txt (#18273) I am happy to remove some nolints for you! --- scripts/nolints.txt | 6 ------ 1 file changed, 6 deletions(-) diff --git a/scripts/nolints.txt b/scripts/nolints.txt index 0da58c64c895a..52ccde41118cc 100644 --- a/scripts/nolints.txt +++ b/scripts/nolints.txt @@ -18,12 +18,6 @@ apply_nolint partrec doc_blame apply_nolint partrec₂ doc_blame -- computability/primrec.lean -apply_nolint nat.cases doc_blame -apply_nolint nat.elim doc_blame -apply_nolint nat.primrec'.vec doc_blame -apply_nolint nat.unpaired doc_blame -apply_nolint primcodable.of_equiv doc_blame -apply_nolint primcodable.subtype doc_blame apply_nolint primrec.dom_fintype fintype_finite -- computability/turing_machine.lean From 6f870fad180eb0f5d8f1e4f385216c25513cb2f8 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Tue, 24 Jan 2023 06:21:12 +0000 Subject: [PATCH 0338/1291] doc(order/game_add): improve docstrings (#18269) We add simpler descriptions for `game_add` and `well_founded.prod_game_add`, without removing previous references to combinatorial games. Mathlib 4 pair: https://github.com/leanprover-community/mathlib4/pull/1798 --- src/order/game_add.lean | 23 ++++++++++++++--------- 1 file changed, 14 insertions(+), 9 deletions(-) diff --git a/src/order/game_add.lean b/src/order/game_add.lean index f7207da171cae..0a9f517cd7ade 100644 --- a/src/order/game_add.lean +++ b/src/order/game_add.lean @@ -33,14 +33,17 @@ variables {α β : Type*} (rα : α → α → Prop) (rβ : β → β → Prop) namespace prod -/-- The "addition of games" relation in combinatorial game theory, on the product type: if - `rα a' a` means that `a ⟶ a'` is a valid move in game `α`, and `rβ b' b` means that `b ⟶ b'` - is a valid move in game `β`, then `game_add rα rβ` specifies the valid moves in the juxtaposition - of `α` and `β`: the player is free to choose one of the games and make a move in it, - while leaving the other game unchanged. -/ +/-- `prod.game_add rα rβ x y` means that `x` can be reached from `y` by decreasing either entry with + respect to the relations `rα` and `rβ`. + + It is so called, as it models game addition within combinatorial game theory. If `rα a₁ a₂` means + that `a₂ ⟶ a₁` is a valid move in game `α`, and `rβ b₁ b₂` means that `b₂ ⟶ b₁` is a valid move + in game `β`, then `game_add rα rβ` specifies the valid moves in the juxtaposition of `α` and `β`: + the player is free to choose one of the games and make a move in it, while leaving the other game + unchanged. -/ inductive game_add : α × β → α × β → Prop -| fst {a' a b} : rα a' a → game_add (a',b) (a,b) -| snd {a b' b} : rβ b' b → game_add (a,b') (a,b) +| fst {a₁ a₂ b} : rα a₁ a₂ → game_add (a₁, b) (a₂, b) +| snd {a b₁ b₂} : rβ b₁ b₂ → game_add (a, b₁) (a, b₂) /-- `game_add` is a `subrelation` of `prod.lex`. -/ lemma game_add_le_lex : game_add rα rβ ≤ prod.lex rα rβ := @@ -69,6 +72,8 @@ begin exacts [iha _ ra (acc.intro b hb), ihb _ rb], end -/-- The sum of two well-founded games is well-founded. -/ +/-- The `prod.game_add` relation on well-founded inputs is well-founded. + + In particular, the sum of two well-founded games is well-founded. -/ lemma well_founded.prod_game_add (hα : well_founded rα) (hβ : well_founded rβ) : - well_founded (prod.game_add rα rβ) := ⟨λ ⟨a,b⟩, (hα.apply a).prod_game_add (hβ.apply b)⟩ + well_founded (prod.game_add rα rβ) := ⟨λ ⟨a, b⟩, (hα.apply a).prod_game_add (hβ.apply b)⟩ From b69c9a770ecf37eb21f7b8cf4fa00de3b62694ec Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Tue, 24 Jan 2023 12:36:48 +0000 Subject: [PATCH 0339/1291] feat(ring_theory/localization): localization of a basis, again (#18261) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR replaces `basis.localization` with `basis.localization_localization`, which maps `basis ι R S` to `basis ι Rm Sm`, where `Rm` and `Sm` are the localizations of `R` and `S` at `m`. This differs from the existing `basis.localization` by localizing in two places at once, since localizing only the coefficients is probably not useful. See also: #18150 Co-authored-by: Vierkantor --- src/ring_theory/localization/basic.lean | 75 ++++++++++++++++++++-- src/ring_theory/localization/module.lean | 82 +++++++++++++++++++++--- 2 files changed, 141 insertions(+), 16 deletions(-) diff --git a/src/ring_theory/localization/basic.lean b/src/ring_theory/localization/basic.lean index 63038c1b0b0f0..f26c26a139e19 100644 --- a/src/ring_theory/localization/basic.lean +++ b/src/ring_theory/localization/basic.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, Mario Carneiro, Johan Commelin, Amelia Livingston, Anne Baanen -/ -import algebra.algebra.equiv +import algebra.algebra.tower import algebra.ring.equiv import group_theory.monoid_localization import ring_theory.ideal.basic @@ -1143,7 +1143,12 @@ variables (S M) Given an algebra `R → S`, a submonoid `R` of `M`, and a localization `Rₘ` for `M`, let `Sₘ` be the localization of `S` to the image of `M` under `algebra_map R S`. Then this is the natural algebra structure on `Rₘ → Sₘ`, such that the entire square commutes, -where `localization_map.map_comp` gives the commutativity of the underlying maps -/ +where `localization_map.map_comp` gives the commutativity of the underlying maps. + +This instance can be helpful if you define `Sₘ := localization (algebra.algebra_map_submonoid S M)`, +however we will instead use the hypotheses `[algebra Rₘ Sₘ] [is_scalar_tower R Rₘ Sₘ]` in lemmas +since the algebra structure may arise in different ways. +-/ noncomputable def localization_algebra : algebra Rₘ Sₘ := (map Sₘ (algebra_map R S) (show _ ≤ (algebra.algebra_map_submonoid S M).comap _, from M.le_comap_map) @@ -1151,10 +1156,68 @@ noncomputable def localization_algebra : algebra Rₘ Sₘ := end -lemma algebra_map_mk' (r : R) (m : M) : - (@algebra_map Rₘ Sₘ _ _ (localization_algebra M S)) (mk' Rₘ r m) = - mk' Sₘ (algebra_map R S r) ⟨algebra_map R S m, algebra.mem_algebra_map_submonoid_of_mem m⟩ := -map_mk' _ _ _ +section + +variables [algebra Rₘ Sₘ] [algebra R Sₘ] [is_scalar_tower R Rₘ Sₘ] [is_scalar_tower R S Sₘ] + +variables (S Rₘ Sₘ) +include S + +lemma is_localization.map_units_map_submonoid (y : M) : is_unit (algebra_map R Sₘ y) := +begin + rw is_scalar_tower.algebra_map_apply _ S, + exact is_localization.map_units Sₘ ⟨algebra_map R S y, algebra.mem_algebra_map_submonoid_of_mem y⟩ +end + +@[simp] lemma is_localization.algebra_map_mk' (x : R) (y : M) : + algebra_map Rₘ Sₘ (is_localization.mk' Rₘ x y) = + is_localization.mk' Sₘ (algebra_map R S x) ⟨algebra_map R S y, + algebra.mem_algebra_map_submonoid_of_mem y⟩ := +begin + rw [is_localization.eq_mk'_iff_mul_eq, subtype.coe_mk, ← is_scalar_tower.algebra_map_apply, + ← is_scalar_tower.algebra_map_apply, is_scalar_tower.algebra_map_apply R Rₘ Sₘ, + is_scalar_tower.algebra_map_apply R Rₘ Sₘ, ← _root_.map_mul, + mul_comm, is_localization.mul_mk'_eq_mk'_of_mul], + exact congr_arg (algebra_map Rₘ Sₘ) (is_localization.mk'_mul_cancel_left x y) +end + +variables (M) + +/-- If the square below commutes, the bottom map is uniquely specified: +``` +R → S +↓ ↓ +Rₘ → Sₘ +``` +-/ +lemma is_localization.algebra_map_eq_map_map_submonoid : + algebra_map Rₘ Sₘ = map Sₘ (algebra_map R S) + (show _ ≤ (algebra.algebra_map_submonoid S M).comap _, from M.le_comap_map) := +eq.symm $ is_localization.map_unique _ (algebra_map Rₘ Sₘ) (λ x, + by rw [← is_scalar_tower.algebra_map_apply R S Sₘ, ← is_scalar_tower.algebra_map_apply R Rₘ Sₘ]) + +/-- If the square below commutes, the bottom map is uniquely specified: +``` +R → S +↓ ↓ +Rₘ → Sₘ +``` +-/ +lemma is_localization.algebra_map_apply_eq_map_map_submonoid (x) : + algebra_map Rₘ Sₘ x = map Sₘ (algebra_map R S) + (show _ ≤ (algebra.algebra_map_submonoid S M).comap _, from M.le_comap_map) + x := +fun_like.congr_fun (is_localization.algebra_map_eq_map_map_submonoid _ _ _ _) x + +variables {R} + +lemma is_localization.lift_algebra_map_eq_algebra_map : + @is_localization.lift R _ M Rₘ _ _ Sₘ _ _ (algebra_map R Sₘ) + (is_localization.map_units_map_submonoid S Sₘ) = + algebra_map Rₘ Sₘ := +is_localization.lift_unique _ (λ x, (is_scalar_tower.algebra_map_apply _ _ _ _).symm) + +end variables (Rₘ Sₘ) diff --git a/src/ring_theory/localization/module.lean b/src/ring_theory/localization/module.lean index 68953652dc66a..b83bb0035a09f 100644 --- a/src/ring_theory/localization/module.lean +++ b/src/ring_theory/localization/module.lean @@ -16,8 +16,8 @@ This file contains some results about vector spaces over the field of fractions * `linear_independent.localization`: `b` is linear independent over a localization of `R` if it is linear independent over `R` itself - * `basis.localization`: promote an `R`-basis `b` to an `Rₛ`-basis, - where `Rₛ` is a localization of `R` + * `basis.localization_localization`: promote an `R`-basis `b` of `A` to an `Rₛ`-basis of `Aₛ`, + where `Rₛ` and `Aₛ` are localizations of `R` and `A` at `s` respectively * `linear_independent.iff_fraction_ring`: `b` is linear independent over `R` iff it is linear independent over `Frac(R)` -/ @@ -54,16 +54,78 @@ begin end end add_comm_monoid -section add_comm_group -variables {M : Type*} [add_comm_group M] [module R M] [module Rₛ M] [is_scalar_tower R Rₛ M] +section localization_localization -/-- Promote a basis for `M` over `R` to a basis for `M` over the localization `Rₛ` -/ -noncomputable def basis.localization {ι : Type*} (b : basis ι R M) : basis ι Rₛ M := -basis.mk (b.linear_independent.localization Rₛ S) $ -by { rw [← eq_top_iff, ← @submodule.restrict_scalars_eq_top_iff Rₛ R, eq_top_iff, ← b.span_eq], - apply submodule.span_le_restrict_scalars } +variables {A : Type*} [comm_ring A] [algebra R A] +variables (Aₛ : Type*) [comm_ring Aₛ] [algebra A Aₛ] +variables [algebra Rₛ Aₛ] [algebra R Aₛ] [is_scalar_tower R Rₛ Aₛ] [is_scalar_tower R A Aₛ] +variables [hA : is_localization (algebra.algebra_map_submonoid A S) Aₛ] +include hA -end add_comm_group +open submodule + +lemma linear_independent.localization_localization {ι : Type*} + {v : ι → A} (hv : linear_independent R v) (hS : algebra.algebra_map_submonoid A S ≤ A⁰) : + linear_independent Rₛ (algebra_map A Aₛ ∘ v) := +begin + refine (hv.map' ((algebra.linear_map A Aₛ).restrict_scalars R) _).localization Rₛ S, + rw [linear_map.ker_restrict_scalars, restrict_scalars_eq_bot_iff, linear_map.ker_eq_bot, + algebra.coe_linear_map], + exact is_localization.injective Aₛ hS +end + +lemma span_eq_top.localization_localization {v : set A} (hv : span R v = ⊤) : + span Rₛ (algebra_map A Aₛ '' v) = ⊤ := +begin + rw eq_top_iff, + rintros a' -, + obtain ⟨a, ⟨_, s, hs, rfl⟩, rfl⟩ := is_localization.mk'_surjective + (algebra.algebra_map_submonoid A S) a', + rw [is_localization.mk'_eq_mul_mk'_one, mul_comm, ← map_one (algebra_map R A)], + erw ← is_localization.algebra_map_mk' A Rₛ Aₛ (1 : R) ⟨s, hs⟩, -- `erw` needed to unify `⟨s, hs⟩` + rw ← algebra.smul_def, + refine smul_mem _ _ (span_subset_span R _ _ _), + rw [← algebra.coe_linear_map, ← linear_map.coe_restrict_scalars R, ← linear_map.map_span], + exact mem_map_of_mem (hv.symm ▸ mem_top), + { apply_instance } +end + +/-- If `A` has an `R`-basis, then localizing `A` at `S` has a basis over `R` localized at `S`. + +A suitable instance for `[algebra A Aₛ]` is `localization_algebra`. +-/ +noncomputable def basis.localization_localization {ι : Type*} (b : basis ι R A) + (hS : algebra.algebra_map_submonoid A S ≤ A⁰) : + basis ι Rₛ Aₛ := +basis.mk + (b.linear_independent.localization_localization _ S _ hS) + (by { rw [set.range_comp, span_eq_top.localization_localization Rₛ S Aₛ b.span_eq], + exact le_rfl }) + +@[simp] lemma basis.localization_localization_apply {ι : Type*} (b : basis ι R A) + (hS : algebra.algebra_map_submonoid A S ≤ A⁰) (i) : + b.localization_localization Rₛ S Aₛ hS i = algebra_map A Aₛ (b i) := +basis.mk_apply _ _ _ + +@[simp] lemma basis.localization_localization_repr_algebra_map {ι : Type*} (b : basis ι R A) + (hS : algebra.algebra_map_submonoid A S ≤ A⁰) (x i) : + (b.localization_localization Rₛ S Aₛ hS).repr (algebra_map A Aₛ x) i = + algebra_map R Rₛ (b.repr x i) := +calc (b.localization_localization Rₛ S Aₛ hS).repr (algebra_map A Aₛ x) i + = (b.localization_localization Rₛ S Aₛ hS).repr + ((b.repr x).sum (λ j c, algebra_map R Rₛ c • algebra_map A Aₛ (b j))) i : + by simp_rw [is_scalar_tower.algebra_map_smul, algebra.smul_def, + is_scalar_tower.algebra_map_apply R A Aₛ, ← _root_.map_mul, ← map_finsupp_sum, + ← algebra.smul_def, ← finsupp.total_apply, basis.total_repr] +... = (b.repr x).sum (λ j c, algebra_map R Rₛ c • finsupp.single j 1 i) : + by simp_rw [← b.localization_localization_apply Rₛ S Aₛ hS, map_finsupp_sum, + linear_equiv.map_smul, basis.repr_self, finsupp.sum_apply, finsupp.smul_apply] +... = _ : finset.sum_eq_single i + (λ j _ hj, by simp [hj]) + (λ hi, by simp [finsupp.not_mem_support_iff.mp hi]) +... = algebra_map R Rₛ (b.repr x i) : by simp [algebra.smul_def] + +end localization_localization end localization From 63f84d91dd847f50bae04a01071f3a5491934e36 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Tue, 24 Jan 2023 14:41:16 +0000 Subject: [PATCH 0340/1291] chore(*): add mathlib4 synchronization comments (#18272) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.big_operators.finprod` * `algebra.group.unique_prods` * `algebra.star.pointwise` * `algebra.tropical.big_operators` * `combinatorics.hall.finite` * `combinatorics.pigeonhole` * `data.countable.basic` * `data.fin.vec_notation` * `data.fintype.order` * `data.typevec` * `data.vector.zip` * `data.zmod.defs` * `logic.equiv.fin` * `order.filter.basic` * `order.filter.curry` * `order.filter.extr` * `order.filter.prod` --- src/algebra/big_operators/finprod.lean | 3 +++ src/algebra/group/unique_prods.lean | 3 +++ src/algebra/star/pointwise.lean | 3 +++ src/algebra/tropical/big_operators.lean | 3 +++ src/combinatorics/hall/finite.lean | 3 +++ src/combinatorics/pigeonhole.lean | 3 +++ src/data/countable/basic.lean | 3 +++ src/data/fin/vec_notation.lean | 3 +++ src/data/fintype/order.lean | 3 +++ src/data/typevec.lean | 3 +++ src/data/vector/zip.lean | 3 +++ src/data/zmod/defs.lean | 3 +++ src/logic/equiv/fin.lean | 3 +++ src/order/filter/basic.lean | 3 +++ src/order/filter/curry.lean | 3 +++ src/order/filter/extr.lean | 3 +++ src/order/filter/prod.lean | 3 +++ 17 files changed, 51 insertions(+) diff --git a/src/algebra/big_operators/finprod.lean b/src/algebra/big_operators/finprod.lean index 1f10a690f35a1..03f5e99ad9b87 100644 --- a/src/algebra/big_operators/finprod.lean +++ b/src/algebra/big_operators/finprod.lean @@ -9,6 +9,9 @@ import algebra.indicator_function /-! # Finite products and sums over types and sets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define products and sums over types and subsets of types, with no finiteness hypotheses. All infinite products and sums are defined to be junk values (i.e. one or zero). This approach is sometimes easier to use than `finset.sum`, diff --git a/src/algebra/group/unique_prods.lean b/src/algebra/group/unique_prods.lean index 7cb21d46d755b..3fc479cc41994 100644 --- a/src/algebra/group/unique_prods.lean +++ b/src/algebra/group/unique_prods.lean @@ -8,6 +8,9 @@ import data.finset.preimage /-! # Unique products and related notions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A group `G` has *unique products* if for any two non-empty finite subsets `A, B ⊂ G`, there is an element `g ∈ A * B` that can be written uniquely as a product of an element of `A` and an element of `B`. We call the formalization this property `unique_prods`. Since the condition requires no diff --git a/src/algebra/star/pointwise.lean b/src/algebra/star/pointwise.lean index 0b27e91d40628..f4b883d973333 100644 --- a/src/algebra/star/pointwise.lean +++ b/src/algebra/star/pointwise.lean @@ -10,6 +10,9 @@ import data.set.pointwise.basic /-! # Pointwise star operation on sets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the star operation pointwise on sets and provides the basic API. Besides basic facts about about how the star operation acts on sets (e.g., `(s ∩ t)⋆ = s⋆ ∩ t⋆`), if `s t : set α`, then under suitable assumption on `α`, it is shown diff --git a/src/algebra/tropical/big_operators.lean b/src/algebra/tropical/big_operators.lean index ac06a7bf8b825..5322f25dfa18b 100644 --- a/src/algebra/tropical/big_operators.lean +++ b/src/algebra/tropical/big_operators.lean @@ -12,6 +12,9 @@ import order.conditionally_complete_lattice.finset # Tropicalization of finitary operations +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file provides the "big-op" or notation-based finitary operations on tropicalized types. This allows easy conversion between sums to Infs and prods to sums. Results here are important for expressing that evaluation of tropical polynomials are the minimum over a finite piecewise diff --git a/src/combinatorics/hall/finite.lean b/src/combinatorics/hall/finite.lean index 921de79ee91db..e28815707c68d 100644 --- a/src/combinatorics/hall/finite.lean +++ b/src/combinatorics/hall/finite.lean @@ -9,6 +9,9 @@ import data.set.finite /-! # Hall's Marriage Theorem for finite index types +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This module proves the basic form of Hall's theorem. In constrast to the theorem described in `combinatorics.hall.basic`, this version requires that the indexed family `t : ι → finset α` have `ι` be finite. diff --git a/src/combinatorics/pigeonhole.lean b/src/combinatorics/pigeonhole.lean index 87bea38b4689a..b16afa592b905 100644 --- a/src/combinatorics/pigeonhole.lean +++ b/src/combinatorics/pigeonhole.lean @@ -12,6 +12,9 @@ import algebra.module.big_operators /-! # Pigeonhole principles +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given pigeons (possibly infinitely many) in pigeonholes, the pigeonhole principle states that, if there are more pigeons than pigeonholes, then there is a pigeonhole with two or more pigeons. diff --git a/src/data/countable/basic.lean b/src/data/countable/basic.lean index 2bee01a4630ba..f167ba762df1a 100644 --- a/src/data/countable/basic.lean +++ b/src/data/countable/basic.lean @@ -10,6 +10,9 @@ import data.countable.defs /-! # Countable types +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we provide basic instances of the `countable` typeclass defined elsewhere. -/ diff --git a/src/data/fin/vec_notation.lean b/src/data/fin/vec_notation.lean index dbc377afeeb7e..29033f48f7911 100644 --- a/src/data/fin/vec_notation.lean +++ b/src/data/fin/vec_notation.lean @@ -11,6 +11,9 @@ import meta.univs /-! # Matrix and vector notation +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines notation for vectors and matrices. Given `a b c d : α`, the notation allows us to write `![a, b, c, d] : fin 4 → α`. Nesting vectors gives coefficients of a matrix, so `![![a, b], ![c, d]] : fin 2 → fin 2 → α`. diff --git a/src/data/fintype/order.lean b/src/data/fintype/order.lean index 0a2de3b661247..b3234143e11f7 100644 --- a/src/data/fintype/order.lean +++ b/src/data/fintype/order.lean @@ -9,6 +9,9 @@ import data.finset.order /-! # Order structures on finite types +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file provides order instances on fintypes. ## Computable instances diff --git a/src/data/typevec.lean b/src/data/typevec.lean index 818c64bb089ff..60416b60ebafb 100644 --- a/src/data/typevec.lean +++ b/src/data/typevec.lean @@ -11,6 +11,9 @@ import tactic.basic # Tuples of types, and their categorical structure. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Features * `typevec n` - n-tuples of types diff --git a/src/data/vector/zip.lean b/src/data/vector/zip.lean index e47bc9ee582ad..001873b548393 100644 --- a/src/data/vector/zip.lean +++ b/src/data/vector/zip.lean @@ -8,6 +8,9 @@ import data.list.zip /-! # The `zip_with` operation on vectors. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ namespace vector diff --git a/src/data/zmod/defs.lean b/src/data/zmod/defs.lean index 3a91f1a2e5f37..a8a6a3732836c 100644 --- a/src/data/zmod/defs.lean +++ b/src/data/zmod/defs.lean @@ -10,6 +10,9 @@ import data.fintype.lattice /-! # Definition of `zmod n` + basic results. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file provides the basic details of `zmod n`, including its commutative ring structure. ## Implementation details diff --git a/src/logic/equiv/fin.lean b/src/logic/equiv/fin.lean index a5b14633c8397..bc227e87d6552 100644 --- a/src/logic/equiv/fin.lean +++ b/src/logic/equiv/fin.lean @@ -8,6 +8,9 @@ import logic.equiv.defs /-! # Equivalences for `fin n` + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ universes u diff --git a/src/order/filter/basic.lean b/src/order/filter/basic.lean index 79f73a8030deb..e98ebb54d07cf 100644 --- a/src/order/filter/basic.lean +++ b/src/order/filter/basic.lean @@ -11,6 +11,9 @@ import tactic.monotonicity /-! # Theory of filters on sets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions * `filter` : filters on a set; diff --git a/src/order/filter/curry.lean b/src/order/filter/curry.lean index d0ad3eee57f57..49208acf1904c 100644 --- a/src/order/filter/curry.lean +++ b/src/order/filter/curry.lean @@ -8,6 +8,9 @@ import order.filter.prod /-! # Curried Filters +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file provides an operation (`filter.curry`) on filters which provides the equivalence `∀ᶠ a in l, ∀ᶠ b in l', p (a, b) ↔ ∀ᶠ c in (l.curry l'), p c` (see `filter.eventually_curry_iff`). diff --git a/src/order/filter/extr.lean b/src/order/filter/extr.lean index 937e748bcbce8..b13ff1f893923 100644 --- a/src/order/filter/extr.lean +++ b/src/order/filter/extr.lean @@ -8,6 +8,9 @@ import order.filter.basic /-! # Minimum and maximum w.r.t. a filter and on a aet +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main Definitions This file defines six predicates of the form `is_A_B`, where `A` is `min`, `max`, or `extr`, diff --git a/src/order/filter/prod.lean b/src/order/filter/prod.lean index 2d8134c16d6b6..b6dc75e02a0d8 100644 --- a/src/order/filter/prod.lean +++ b/src/order/filter/prod.lean @@ -8,6 +8,9 @@ import order.filter.basic /-! # Product and coproduct filters +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define `filter.prod f g` (notation: `f ×ᶠ g`) and `filter.coprod f g`. The product of two filters is the largest filter `l` such that `filter.tendsto prod.fst l f` and `filter.tendsto prod.snd l g`. From e3d9ab8faa9dea8f78155c6c27d62a621f4c152d Mon Sep 17 00:00:00 2001 From: Floris van Doorn Date: Tue, 24 Jan 2023 17:31:34 +0000 Subject: [PATCH 0341/1291] feat(tactic/simps): store generated lemmas in attribute (#17964) * `@[simps] def equiv.prod` now adds an attribute ``@[_simps_aux [`equiv.prod_apply, `equiv.prod_symm_apply]]`` to `equiv.prod` * Remark: the lemmas generated by `simps` can have `to_additive` attributes themselves, if the original declaration had a `to_additive` attribute (but it doesn't recurse any deeper than that) * Can be used by mathport to generate `#align` statements --- src/tactic/simps.lean | 65 +++++++++++++++++++++++++++---------------- test/simps.lean | 11 ++++++-- 2 files changed, 49 insertions(+), 27 deletions(-) diff --git a/src/tactic/simps.lean b/src/tactic/simps.lean index fc4a67b1a725a..b66df7fc71e1c 100644 --- a/src/tactic/simps.lean +++ b/src/tactic/simps.lean @@ -560,7 +560,7 @@ meta def simps_get_projection_exprs (e : environment) (tgt : expr) /-- Add a lemma with `nm` stating that `lhs = rhs`. `type` is the type of both `lhs` and `rhs`, `args` is the list of local constants occurring, and `univs` is the list of universe variables. -/ meta def simps_add_projection (nm : name) (type lhs rhs : expr) (args : list expr) - (univs : list name) (cfg : simps_cfg) : tactic unit := do + (univs : list name) (cfg : simps_cfg) : tactic (list name) := do when_tracing `simps.debug trace! "[simps] > Planning to add the equality\n > {lhs} = ({rhs} : {type})", lvl ← get_univ_level type, @@ -587,7 +587,8 @@ meta def simps_add_projection (nm : name) (type lhs rhs : expr) (args : list exp when (b ∧ `simp ∈ cfg.attrs) (set_basic_attribute `_refl_lemma decl_name tt), cfg.attrs.mmap' $ λ nm, set_attribute nm decl_name tt, when cfg.add_additive.is_some $ - to_additive.attr.set decl_name ⟨ff, cfg.trace, cfg.add_additive.iget, none, tt⟩ tt + to_additive.attr.set decl_name ⟨ff, cfg.trace, cfg.add_additive.iget, none, tt⟩ tt, + pure [decl_name] /-- Derive lemmas specifying the projections of the declaration. If `todo` is non-empty, it will generate exactly the names in `todo`. @@ -595,7 +596,7 @@ meta def simps_add_projection (nm : name) (type lhs rhs : expr) (args : list exp was just used. In that case we need to apply these projections before we continue changing lhs. -/ meta def simps_add_projections : Π (e : environment) (nm : name) (type lhs rhs : expr) (args : list expr) (univs : list name) (must_be_str : bool) - (cfg : simps_cfg) (todo : list string) (to_apply : list ℕ), tactic unit + (cfg : simps_cfg) (todo : list string) (to_apply : list ℕ), tactic (list name) | e nm type lhs rhs args univs must_be_str cfg todo to_apply := do -- we don't want to unfold non-reducible definitions (like `set`) to apply more arguments when_tracing `simps.debug trace! @@ -615,18 +616,19 @@ meta def simps_add_projections : Π (e : environment) (nm : name) if e.is_structure str ∧ ¬(todo = [] ∧ str ∈ cfg.not_recursive ∧ ¬must_be_str) then do [intro] ← return $ e.constructors_of str | fail "unreachable code (3)", rhs_whnf ← whnf rhs_ap cfg.rhs_md, - (rhs_ap, todo_now) ← -- `todo_now` means that we still have to generate the current simp lemma + -- `todo_now` means that we still have to generate the current simp lemma + (rhs_ap, todo_now, added_lems_requested) ← if ¬ is_constant_of rhs_ap.get_app_fn intro ∧ - is_constant_of rhs_whnf.get_app_fn intro then + is_constant_of rhs_whnf.get_app_fn intro then do /- If this was a desired projection, we want to apply it before taking the whnf. However, if the current field is an eta-expansion (see below), we first want to eta-reduce it and only then construct the projection. This makes the flow of this function messy. -/ - when ("" ∈ todo ∧ to_apply = []) (if cfg.fully_applied then + added_lems_requested ← cond ("" ∈ todo ∧ to_apply = []) (if cfg.fully_applied then simps_add_projection nm tgt lhs_ap rhs_ap new_args univs cfg else - simps_add_projection nm type lhs rhs args univs cfg) >> - return (rhs_whnf, ff) else - return (rhs_ap, "" ∈ todo ∧ to_apply = []), + simps_add_projection nm type lhs rhs args univs cfg) (pure []), + return (rhs_whnf, ff, added_lems_requested) else + return (rhs_ap, "" ∈ todo ∧ to_apply = [], []), if is_constant_of (get_app_fn rhs_ap) intro then do -- if the value is a constructor application proj_info ← simps_get_projection_exprs e tgt rhs_ap cfg, when_tracing `simps.debug trace!"[simps] > Raw projection information:\n {proj_info}", @@ -635,12 +637,12 @@ meta def simps_add_projections : Π (e : environment) (nm : name) /- As a special case, we want to automatically generate the current projection if `rhs_ap` was an eta-expansion. Also, when this was a desired projection, we need to generate the current projection if we haven't done it above. -/ - when (todo_now ∨ (todo = [] ∧ eta.is_some ∧ to_apply = [])) $ - if cfg.fully_applied then + added_lems_eta ← cond (todo_now ∨ (todo = [] ∧ eta.is_some ∧ to_apply = [])) + (if cfg.fully_applied then simps_add_projection nm tgt lhs_ap rhs_ap new_args univs cfg else - simps_add_projection nm type lhs rhs args univs cfg, + simps_add_projection nm type lhs rhs args univs cfg) (return []), /- If we are in the middle of a composite projection. -/ - when (to_apply ≠ []) $ do + added_lems_custom_proj ← cond (to_apply ≠ []) (do { ⟨new_rhs, proj, proj_expr, proj_nrs, is_default, is_prefix⟩ ← return $ proj_info.inth to_apply.head, new_type ← infer_type new_rhs, @@ -648,10 +650,11 @@ meta def simps_add_projections : Π (e : environment) (nm : name) trace!"[simps] > Applying a custom composite projection. Current lhs: > {lhs_ap}", simps_add_projections e nm new_type lhs_ap new_rhs new_args univs ff cfg todo - to_apply.tail }, + to_apply.tail }) (pure []), + let all_added_lems := added_lems_requested ++ added_lems_eta ++ added_lems_custom_proj, /- We stop if no further projection is specified or if we just reduced an eta-expansion and we automatically choose projections -/ - when ¬(to_apply ≠ [] ∨ todo = [""] ∨ (eta.is_some ∧ todo = [])) $ do + cond (¬(to_apply ≠ [] ∨ todo = [""] ∨ (eta.is_some ∧ todo = []))) (do let projs : list name := proj_info.map $ λ x, x.snd.name, let todo := if to_apply = [] then todo_next else todo, -- check whether all elements in `todo` have a projection as prefix @@ -666,13 +669,13 @@ The known projections are: You can also see this information by running `initialize_simps_projections? {str}`. Note: these projection names might not correspond to the projection names of the structure.", - proj_info.mmap_with_index' $ - λ proj_nr ⟨new_rhs, proj, proj_expr, proj_nrs, is_default, is_prefix⟩, do + added_lems_list ← proj_info.mmap_with_index + (λ proj_nr ⟨new_rhs, proj, proj_expr, proj_nrs, is_default, is_prefix⟩, do new_type ← infer_type new_rhs, let new_todo := todo.filter_map $ λ x, x.get_rest ("_" ++ proj.last), -- we only continue with this field if it is non-propositional or mentioned in todo - when ((is_default ∧ todo = []) ∨ new_todo ≠ []) $ do + cond ((is_default ∧ todo = []) ∨ new_todo ≠ []) (do let new_lhs := proj_expr.instantiate_lambdas_or_apps [lhs_ap], let new_nm := nm.append_to_last proj.last is_prefix, let new_cfg := { add_additive := cfg.add_additive.map $ @@ -680,15 +683,17 @@ Note: these projection names might not correspond to the projection names of the when_tracing `simps.debug trace!"[simps] > Recursively add projections for: > {new_lhs}", simps_add_projections e new_nm new_type new_lhs new_rhs new_args univs - ff new_cfg new_todo proj_nrs + ff new_cfg new_todo proj_nrs) (pure [])), + pure $ all_added_lems ++ added_lems_list.join) (pure all_added_lems) -- if I'm about to run into an error, try to set the transparency for `rhs_md` higher. else if cfg.rhs_md = transparency.none ∧ (must_be_str ∨ todo_next ≠ [] ∨ to_apply ≠ []) then do when cfg.trace trace! "[simps] > The given definition is not a constructor application: > {rhs_ap} > Retrying with the options {{ rhs_md := semireducible, simp_rhs := tt}.", - simps_add_projections e nm type lhs rhs args univs must_be_str - { rhs_md := semireducible, simp_rhs := tt, ..cfg} todo to_apply + added_lems_recursive ← simps_add_projections e nm type lhs rhs args univs must_be_str + { rhs_md := semireducible, simp_rhs := tt, ..cfg} todo to_apply, + pure $ added_lems_requested ++ added_lems_recursive else do when (to_apply ≠ []) $ fail!"Invalid simp lemma {nm}. @@ -698,9 +703,10 @@ The given definition is not a constructor application:\n {rhs_ap}", when (todo_next ≠ []) $ fail!"Invalid simp lemma {nm.append_suffix todo_next.head}. The given definition is not a constructor application:\n {rhs_ap}", - if cfg.fully_applied then + added_lems_no_constructor ← if cfg.fully_applied then simps_add_projection nm tgt lhs_ap rhs_ap new_args univs cfg else - simps_add_projection nm type lhs rhs args univs cfg + simps_add_projection nm type lhs rhs args univs cfg, + pure $ added_lems_requested ++ added_lems_no_constructor else do when must_be_str $ fail!"Invalid `simps` attribute. Target {str} is not a structure", @@ -712,6 +718,15 @@ Projection {(first_todo.split_on '_').tail.head} doesn't exist, because target i simps_add_projection nm tgt lhs_ap rhs_ap new_args univs cfg else simps_add_projection nm type lhs rhs args univs cfg +/-- +The `@[_simps_aux]` attribute specifies which lemmas are added by `simps`. +This should not be used manually and it only exists for mathport +-/ +@[user_attribute] meta def simps_aux : user_attribute unit (list name) := +{ name := `_simps_aux, + descr := "An attribute specifying the added simps lemmas.", + parser := failed } + /-- `simps_tac` derives `simp` lemmas for all (nested) non-Prop projections of the declaration. If `todo` is non-empty, it will generate exactly the names in `todo`. If `short_nm` is true, the generated names will only use the last projection name. @@ -730,7 +745,9 @@ meta def simps_tac (nm : name) (cfg : simps_cfg := {}) (todo : list string := [] trace!"[simps] > @[to_additive] will be added to all generated lemmas.", return { add_additive := dict.find nm, ..cfg } } else return cfg, - simps_add_projections e nm d.type lhs d.value [] d.univ_params tt cfg todo [] + added_names ← simps_add_projections e nm d.type lhs d.value [] d.univ_params tt cfg todo [], + simps_aux.set nm added_names true + /-- The parser for the `@[simps]` attribute. -/ meta def simps_parser : parser (bool × list string × simps_cfg) := do diff --git a/test/simps.lean b/test/simps.lean index 74a558af4b53f..665f3c16082e3 100644 --- a/test/simps.lean +++ b/test/simps.lean @@ -34,7 +34,9 @@ run_cmd do e.get `foo.rfl_to_fun, e.get `foo.rfl_inv_fun, success_if_fail (e.get `foo.rfl_left_inv), - success_if_fail (e.get `foo.rfl_right_inv) + success_if_fail (e.get `foo.rfl_right_inv), + p ← simps_aux.get_param `foo.rfl, + guard $ p = [`foo.rfl_to_fun, `foo.rfl_inv_fun] example (n : ℕ) : foo.rfl.to_fun n = n := by rw [foo.rfl_to_fun, id] example (n : ℕ) : foo.rfl.inv_fun n = n := by rw [foo.rfl_inv_fun] @@ -256,7 +258,9 @@ Note: these projection names might not correspond to the projection names of the success_if_fail_with_msg (simps_tac `specify.specify5 {} ["snd_snd"]) "Invalid simp lemma specify.specify5_snd_snd. The given definition is not a constructor application: - classical.choice specify.specify5._proof_1" + classical.choice specify.specify5._proof_1", + p ← simps_aux.get_param `specify.specify4, + guard $ p = [`specify.specify4_snd, `specify.specify4_snd_snd] /- We also eta-reduce if we explicitly specify the projection. -/ @@ -514,7 +518,8 @@ variables {α β γ : Sort*} /-- See Note [custom simps projection] -/ noncomputable def equiv.simps.inv_fun (e : α ≃ β) : β → α := classical.choice ⟨e.inv_fun⟩ -run_cmd do e ← get_env, success_if_fail_with_msg (simps_get_raw_projections e `faulty_manual_coercion.equiv) +run_cmd do e ← get_env, success_if_fail_with_msg + (simps_get_raw_projections e `faulty_manual_coercion.equiv) "Invalid custom projection: λ {α : Sort u_1} {β : Sort u_2} (e : α ≃ β), classical.choice _ Expression is not definitionally equal to From bbd198e1a40c86d23dd49949c2761216f7ebea75 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 24 Jan 2023 20:54:47 +0000 Subject: [PATCH 0342/1291] chore(test/gmonoid): Add a test of the gmonoid API (#18197) This is taken from my CICM 2022 paper. I don't think the instance is canonical enough to be in mathlib, but it's useful to have a test that verifies it can be constructed. --- test/gmonoid.lean | 30 ++++++++++++++++++++++++++++++ 1 file changed, 30 insertions(+) create mode 100644 test/gmonoid.lean diff --git a/test/gmonoid.lean b/test/gmonoid.lean new file mode 100644 index 0000000000000..0c6c7a8b25eb1 --- /dev/null +++ b/test/gmonoid.lean @@ -0,0 +1,30 @@ +/- +Copyright (c) 2021 Eric Wieser. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Eric Wieser +-/ +import algebra.direct_sum.ring +import data.fin.tuple.basic + +/-! # Tuples `fin n → α` form a graded monoid with `*` as `fin.append` + +Defining multiplication as concatenation isn't particularly canonical, so we do not provide +this in mathlib. We could safely provide this instance on a type alias, but for now we just put +this in `tests` to verify that this definition is possible. -/ + +namespace fin + +variables {α : Type*} {α' : Type*} {na nb nc : ℕ} + +example {α : Type*} : graded_monoid.gmonoid (λ n, fin n → α) := +{ mul := λ i j, fin.append, + one := fin.elim0, + one_mul := λ b, sigma_eq_of_eq_comp_cast _ (elim0'_append _), + mul_one := λ a, sigma_eq_of_eq_comp_cast _ (append_elim0' _), + mul_assoc := λ a b c, + sigma_eq_of_eq_comp_cast (add_assoc _ _ _) $ (append_assoc a.2 b.2 c.2).trans rfl, + gnpow := λ n i a, repeat n a, + gnpow_zero' := λ a, sigma_eq_of_eq_comp_cast _ (repeat_zero _), + gnpow_succ' := λ a n, sigma_eq_of_eq_comp_cast _ (repeat_succ _ _) } + +end fin From a81201cc7a065e7ebda308c7121d39c34cb4f33b Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Tue, 24 Jan 2023 20:54:48 +0000 Subject: [PATCH 0343/1291] feat(topology/algebra): add converse to tendsto.mul (#18276) In a topological group-with-zero, if`f * g` tends to `x * y` and `g` tends to `y`, with `y \ne 0`, then `f` tends to `x`. --- src/topology/algebra/group_with_zero.lean | 11 +++++++++++ 1 file changed, 11 insertions(+) diff --git a/src/topology/algebra/group_with_zero.lean b/src/topology/algebra/group_with_zero.lean index b8f1d99104595..10f1d3ef63ad8 100644 --- a/src/topology/algebra/group_with_zero.lean +++ b/src/topology/algebra/group_with_zero.lean @@ -142,6 +142,17 @@ lemma filter.tendsto.div {l : filter α} {a b : G₀} (hf : tendsto f l (𝓝 a) tendsto (f / g) l (𝓝 (a / b)) := by simpa only [div_eq_mul_inv] using hf.mul (hg.inv₀ hy) +lemma filter.tendsto_mul_iff_of_ne_zero [t1_space G₀] + {f g : α → G₀} {l : filter α} {x y : G₀} + (hg : tendsto g l (𝓝 y)) (hy : y ≠ 0) : + tendsto (λ n, f n * g n) l (𝓝 $ x * y) ↔ tendsto f l (𝓝 x) := +begin + refine ⟨λ hfg, _, λ hf, hf.mul hg⟩, + rw ←mul_div_cancel x hy, + refine tendsto.congr' _ (hfg.div hg hy), + refine eventually.mp (hg.eventually_ne hy) (eventually_of_forall (λ n hn, mul_div_cancel _ hn)), +end + variables [topological_space α] [topological_space β] {s : set α} {a : α} lemma continuous_within_at.div (hf : continuous_within_at f s a) (hg : continuous_within_at g s a) From 6623e6af705e97002a9054c1c05a980180276fc1 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 24 Jan 2023 23:10:50 +0000 Subject: [PATCH 0344/1291] fix(*): add missing `classical` tactics and `decidable` arguments (#18277) As discussed in [this Zulip thread](https://leanprover.zulipchat.com/#narrow/stream/113488-general/topic/.60classical.60.20attribute.20leakage/near/282726128), the `classical` tactic is buggy in Mathlib3, and "leaks" into subsequent declarations. This doesn't port well, as the bug is fixed in lean 4. This PR installs a temporary hack to contain these leaks, fixes all of the correponding breakages, then reverts the hack. The result is that the new `classical` tactics in the diff are not needed in Lean 3, but will be needed in Lean 4. In a future PR, I will try committing the hack itself; but in the meantime, these files are very close to (if not beyond) the port, so the sooner they are fixed the better. --- src/algebra/direct_sum/basic.lean | 10 +++++--- src/algebra/direct_sum/module.lean | 10 +++++--- src/algebra/free_algebra.lean | 4 +-- src/algebra/module/pid.lean | 1 + src/algebra/monoid_algebra/basic.lean | 7 +++--- .../elliptic_curve/weierstrass.lean | 11 ++++++-- .../inner_product_space/orientation.lean | 16 ++++++------ src/analysis/inner_product_space/pi_L2.lean | 6 +++-- src/combinatorics/simple_graph/basic.lean | 9 ++++--- src/data/dfinsupp/basic.lean | 25 +++++++++++++------ src/data/finsupp/alist.lean | 1 + src/data/finsupp/basic.lean | 2 +- src/data/finsupp/defs.lean | 23 +++++++++-------- src/data/finsupp/to_dfinsupp.lean | 1 + src/data/fintype/basic.lean | 2 +- src/data/fintype/sum.lean | 6 ++--- src/data/pi/lex.lean | 7 ++++-- src/group_theory/divisible.lean | 2 +- .../free_module/ideal_quotient.lean | 2 +- .../legendre_symbol/mul_character.lean | 4 ++- src/order/interval.lean | 8 +++--- src/order/max.lean | 1 + src/order/order_iso_nat.lean | 14 +++++++---- src/order/prop_instances.lean | 2 ++ src/ring_theory/bezout.lean | 3 ++- src/ring_theory/discriminant.lean | 1 + src/set_theory/game/basic.lean | 7 +++--- 27 files changed, 121 insertions(+), 64 deletions(-) diff --git a/src/algebra/direct_sum/basic.lean b/src/algebra/direct_sum/basic.lean index 3ebb86c9613d4..0b14312f4daea 100644 --- a/src/algebra/direct_sum/basic.lean +++ b/src/algebra/direct_sum/basic.lean @@ -230,16 +230,20 @@ noncomputable def sigma_curry : (⨁ (i : Σ i, _), δ i.1 i.2) →+ ⨁ i j, δ /--The natural map between `⨁ i (j : α i), δ i j` and `Π₀ (i : Σ i, α i), δ i.1 i.2`, inverse of `curry`.-/ -noncomputable def sigma_uncurry : (⨁ i j, δ i j) →+ ⨁ (i : Σ i, _), δ i.1 i.2 := +def sigma_uncurry [Π i, decidable_eq (α i)] [Π i j, decidable_eq (δ i j)] : + (⨁ i j, δ i j) →+ ⨁ (i : Σ i, _), δ i.1 i.2 := { to_fun := dfinsupp.sigma_uncurry, map_zero' := dfinsupp.sigma_uncurry_zero, map_add' := dfinsupp.sigma_uncurry_add } -@[simp] lemma sigma_uncurry_apply (f : ⨁ i j, δ i j) (i : ι) (j : α i) : +@[simp] lemma sigma_uncurry_apply [Π i, decidable_eq (α i)] [Π i j, decidable_eq (δ i j)] + (f : ⨁ i j, δ i j) (i : ι) (j : α i) : sigma_uncurry f ⟨i, j⟩ = f i j := dfinsupp.sigma_uncurry_apply f i j /--The natural map between `⨁ (i : Σ i, α i), δ i.1 i.2` and `⨁ i (j : α i), δ i j`.-/ -noncomputable def sigma_curry_equiv : (⨁ (i : Σ i, _), δ i.1 i.2) ≃+ ⨁ i j, δ i j := +noncomputable def sigma_curry_equiv + [Π i, decidable_eq (α i)] [Π i j, decidable_eq (δ i j)] : + (⨁ (i : Σ i, _), δ i.1 i.2) ≃+ ⨁ i j, δ i j := { ..sigma_curry, ..dfinsupp.sigma_curry_equiv } end sigma diff --git a/src/algebra/direct_sum/module.lean b/src/algebra/direct_sum/module.lean index a6c323a53a35f..560892bc2a36d 100644 --- a/src/algebra/direct_sum/module.lean +++ b/src/algebra/direct_sum/module.lean @@ -222,15 +222,19 @@ noncomputable def sigma_lcurry : (⨁ (i : Σ i, _), δ i.1 i.2) →ₗ[R] ⨁ i sigma_lcurry R f i j = f ⟨i, j⟩ := sigma_curry_apply f i j /--`uncurry` as a linear map.-/ -noncomputable def sigma_luncurry : (⨁ i j, δ i j) →ₗ[R] ⨁ (i : Σ i, _), δ i.1 i.2 := +def sigma_luncurry [Π i, decidable_eq (α i)] [Π i j, decidable_eq (δ i j)] : + (⨁ i j, δ i j) →ₗ[R] ⨁ (i : Σ i, _), δ i.1 i.2 := { map_smul' := dfinsupp.sigma_uncurry_smul, ..sigma_uncurry } -@[simp] lemma sigma_luncurry_apply (f : ⨁ i j, δ i j) (i : ι) (j : α i) : +@[simp] lemma sigma_luncurry_apply [Π i, decidable_eq (α i)] [Π i j, decidable_eq (δ i j)] + (f : ⨁ i j, δ i j) (i : ι) (j : α i) : sigma_luncurry R f ⟨i, j⟩ = f i j := sigma_uncurry_apply f i j /--`curry_equiv` as a linear equiv.-/ -noncomputable def sigma_lcurry_equiv : (⨁ (i : Σ i, _), δ i.1 i.2) ≃ₗ[R] ⨁ i j, δ i j := +noncomputable def sigma_lcurry_equiv + [Π i, decidable_eq (α i)] [Π i j, decidable_eq (δ i j)] : + (⨁ (i : Σ i, _), δ i.1 i.2) ≃ₗ[R] ⨁ i j, δ i j := { ..sigma_curry_equiv, ..sigma_lcurry R } end sigma diff --git a/src/algebra/free_algebra.lean b/src/algebra/free_algebra.lean index 636e5dbd452f4..965190730dab2 100644 --- a/src/algebra/free_algebra.lean +++ b/src/algebra/free_algebra.lean @@ -344,9 +344,9 @@ map_eq_one_iff (algebra_map _ _) algebra_map_left_inverse.injective -- this proof is copied from the approach in `free_abelian_group.of_injective` lemma ι_injective [nontrivial R] : function.injective (ι R : X → free_algebra R X) := -λ x y hoxy, classical.by_contradiction $ assume hxy : x ≠ y, +λ x y hoxy, classical.by_contradiction $ by classical; exact assume hxy : x ≠ y, let f : free_algebra R X →ₐ[R] R := - lift R (λ z, by classical; exact if x = z then (1 : R) else 0) in + lift R (λ z, if x = z then (1 : R) else 0) in have hfx1 : f (ι R x) = 1, from (lift_ι_apply _ _).trans $ if_pos rfl, have hfy1 : f (ι R y) = 1, from hoxy ▸ hfx1, have hfy0 : f (ι R y) = 0, from (lift_ι_apply _ _).trans $ if_neg hxy, diff --git a/src/algebra/module/pid.lean b/src/algebra/module/pid.lean index 4b3a7fc1c6440..ffdde65fe86c2 100644 --- a/src/algebra/module/pid.lean +++ b/src/algebra/module/pid.lean @@ -213,6 +213,7 @@ begin haveI := λ i, is_noetherian_submodule' (torsion_by R N $ p i ^ e i), exact λ i, torsion_by_prime_power_decomposition (hp i) ((is_torsion'_powers_iff $ p i).mpr $ λ x, ⟨e i, smul_torsion_by _ _⟩) }, + classical, refine ⟨Σ i, fin (this i).some, infer_instance, λ ⟨i, j⟩, p i, λ ⟨i, j⟩, hp i, λ ⟨i, j⟩, (this i).some_spec.some j, ⟨(linear_equiv.of_bijective (direct_sum.coe_linear_map _) h).symm.trans $ diff --git a/src/algebra/monoid_algebra/basic.lean b/src/algebra/monoid_algebra/basic.lean index a09ea8d935603..6df0621c49e15 100644 --- a/src/algebra/monoid_algebra/basic.lean +++ b/src/algebra/monoid_algebra/basic.lean @@ -330,7 +330,8 @@ end lemma mul_apply_antidiagonal [has_mul G] (f g : monoid_algebra k G) (x : G) (s : finset (G × G)) (hs : ∀ {p : G × G}, p ∈ s ↔ p.1 * p.2 = x) : (f * g) x = ∑ p in s, (f p.1 * g p.2) := -let F : G × G → k := λ p, by classical; exact if p.1 * p.2 = x then f p.1 * g p.2 else 0 in +by classical; exact +let F : G × G → k := λ p, if p.1 * p.2 = x then f p.1 * g p.2 else 0 in calc (f * g) x = (∑ a₁ in f.support, ∑ a₂ in g.support, F (a₁, a₂)) : mul_apply f g x ... = ∑ p in f.support ×ˢ g.support, F p : finset.sum_product.symm @@ -477,8 +478,8 @@ end⟩ also commute with the algebra multiplication. -/ instance smul_comm_class_self [smul_comm_class R k k] : smul_comm_class R (monoid_algebra k G) (monoid_algebra k G) := -⟨λ t a b, -begin +⟨λ t a b, begin + classical, ext m, simp only [mul_apply, finsupp.sum, finset.smul_sum, smul_ite, mul_smul_comm, sum_smul_index', implies_true_iff, eq_self_iff_true, coe_smul, ite_eq_right_iff, smul_eq_mul, pi.smul_apply, diff --git a/src/algebraic_geometry/elliptic_curve/weierstrass.lean b/src/algebraic_geometry/elliptic_curve/weierstrass.lean index 791d84b50b381..704fd430cd765 100644 --- a/src/algebraic_geometry/elliptic_curve/weierstrass.lean +++ b/src/algebraic_geometry/elliptic_curve/weierstrass.lean @@ -433,6 +433,8 @@ instance : is_scalar_tower R R[X] W.coordinate_ring := ideal.quotient.is_scalar_ instance [subsingleton R] : subsingleton W.coordinate_ring := module.subsingleton R[X] _ +section +open_locale classical /-- The basis $\{1, Y\}$ for the coordinate ring $R[W]$ over the polynomial ring $R[X]$. Given a Weierstrass curve `W`, write `W^.coordinate_ring.basis` for this basis. -/ @@ -440,11 +442,16 @@ protected noncomputable def basis : basis (fin 2) R[X] W.coordinate_ring := (subsingleton_or_nontrivial R).by_cases (λ _, by exactI default) $ λ _, by exactI (basis.reindex (adjoin_root.power_basis' W.monic_polynomial).basis $ fin_congr $ W.nat_degree_polynomial) +end lemma basis_apply (n : fin 2) : W^.coordinate_ring.basis n = (adjoin_root.power_basis' W.monic_polynomial).gen ^ (n : ℕ) := -by { nontriviality R, simpa only [coordinate_ring.basis, or.by_cases, dif_neg (not_subsingleton R), - basis.reindex_apply, power_basis.basis_eq_pow] } +begin + classical, + nontriviality R, + simpa only [coordinate_ring.basis, or.by_cases, dif_neg (not_subsingleton R), + basis.reindex_apply, power_basis.basis_eq_pow] +end lemma basis_zero : W^.coordinate_ring.basis 0 = 1 := by simpa only [basis_apply] using pow_zero _ diff --git a/src/analysis/inner_product_space/orientation.lean b/src/analysis/inner_product_space/orientation.lean index 605f4775837ae..17a6ec884d5b9 100644 --- a/src/analysis/inner_product_space/orientation.lean +++ b/src/analysis/inner_product_space/orientation.lean @@ -213,7 +213,8 @@ lemma volume_form_robust (b : orthonormal_basis (fin n) ℝ E) (hb : b.to_basis. o.volume_form = b.to_basis.det := begin unfreezingI { cases n }, - { have : o = positive_orientation := hb.symm.trans b.to_basis.orientation_is_empty, + { classical, + have : o = positive_orientation := hb.symm.trans b.to_basis.orientation_is_empty, simp [volume_form, or.by_cases, dif_pos this] }, { dsimp [volume_form], rw [same_orientation_iff_det_eq_det, hb], @@ -227,7 +228,8 @@ lemma volume_form_robust_neg (b : orthonormal_basis (fin n) ℝ E) o.volume_form = - b.to_basis.det := begin unfreezingI { cases n }, - { have : positive_orientation ≠ o := by rwa b.to_basis.orientation_is_empty at hb, + { classical, + have : positive_orientation ≠ o := by rwa b.to_basis.orientation_is_empty at hb, simp [volume_form, or.by_cases, dif_neg this.symm] }, let e : orthonormal_basis (fin n.succ) ℝ E := o.fin_orthonormal_basis n.succ_pos (fact.out _), dsimp [volume_form], @@ -239,7 +241,7 @@ end @[simp] lemma volume_form_neg_orientation : (-o).volume_form = - o.volume_form := begin unfreezingI { cases n }, - { refine o.eq_or_eq_neg_of_is_empty.by_cases _ _; rintros rfl; simp [volume_form_zero_neg] }, + { refine o.eq_or_eq_neg_of_is_empty.elim _ _; rintros rfl; simp [volume_form_zero_neg] }, let e : orthonormal_basis (fin n.succ) ℝ E := o.fin_orthonormal_basis n.succ_pos (fact.out _), have h₁ : e.to_basis.orientation = o := o.fin_orthonormal_basis_orientation _ _, have h₂ : e.to_basis.orientation ≠ -o, @@ -252,7 +254,7 @@ lemma volume_form_robust' (b : orthonormal_basis (fin n) ℝ E) (v : fin n → E |o.volume_form v| = |b.to_basis.det v| := begin unfreezingI { cases n }, - { refine o.eq_or_eq_neg_of_is_empty.by_cases _ _; rintros rfl; simp }, + { refine o.eq_or_eq_neg_of_is_empty.elim _ _; rintros rfl; simp }, { rw [o.volume_form_robust (b.adjust_to_orientation o) (b.orientation_adjust_to_orientation o), b.abs_det_adjust_to_orientation] }, end @@ -263,7 +265,7 @@ value by the product of the norms of the vectors `v i`. -/ lemma abs_volume_form_apply_le (v : fin n → E) : |o.volume_form v| ≤ ∏ i : fin n, ‖v i‖ := begin unfreezingI { cases n }, - { refine o.eq_or_eq_neg_of_is_empty.by_cases _ _; rintros rfl; simp }, + { refine o.eq_or_eq_neg_of_is_empty.elim _ _; rintros rfl; simp }, haveI : finite_dimensional ℝ E := fact_finite_dimensional_of_finrank_eq_succ n, have : finrank ℝ E = fintype.card (fin n.succ) := by simpa using _i.out, let b : orthonormal_basis (fin n.succ) ℝ E := gram_schmidt_orthonormal_basis this v, @@ -288,7 +290,7 @@ lemma abs_volume_form_apply_of_pairwise_orthogonal |o.volume_form v| = ∏ i : fin n, ‖v i‖ := begin unfreezingI { cases n }, - { refine o.eq_or_eq_neg_of_is_empty.by_cases _ _; rintros rfl; simp }, + { refine o.eq_or_eq_neg_of_is_empty.elim _ _; rintros rfl; simp }, haveI : finite_dimensional ℝ E := fact_finite_dimensional_of_finrank_eq_succ n, have hdim : finrank ℝ E = fintype.card (fin n.succ) := by simpa using _i.out, let b : orthonormal_basis (fin n.succ) ℝ E := gram_schmidt_orthonormal_basis hdim v, @@ -320,7 +322,7 @@ lemma volume_form_map {F : Type*} [inner_product_space ℝ F] [fact (finrank ℝ (orientation.map (fin n) φ.to_linear_equiv o).volume_form x = o.volume_form (φ.symm ∘ x) := begin unfreezingI { cases n }, - { refine o.eq_or_eq_neg_of_is_empty.by_cases _ _; rintros rfl; simp }, + { refine o.eq_or_eq_neg_of_is_empty.elim _ _; rintros rfl; simp }, let e : orthonormal_basis (fin n.succ) ℝ E := o.fin_orthonormal_basis n.succ_pos (fact.out _), have he : e.to_basis.orientation = o := (o.fin_orthonormal_basis_orientation n.succ_pos (fact.out _)), diff --git a/src/analysis/inner_product_space/pi_L2.lean b/src/analysis/inner_product_space/pi_L2.lean index d4ae133a5b30f..8449d1837578a 100644 --- a/src/analysis/inner_product_space/pi_L2.lean +++ b/src/analysis/inner_product_space/pi_L2.lean @@ -296,6 +296,7 @@ basis.of_equiv_fun b.repr.to_linear_equiv begin change ⇑(basis.of_equiv_fun b.repr.to_linear_equiv) = b, ext j, + classical, rw basis.coe_of_equiv_fun, congr, end @@ -402,7 +403,7 @@ protected lemma coe_mk (hon : orthonormal 𝕜 v) (hsp: ⊤ ≤ submodule.span by classical; rw [orthonormal_basis.mk, _root_.basis.coe_to_orthonormal_basis, basis.coe_mk] /-- Any finite subset of a orthonormal family is an `orthonormal_basis` for its span. -/ -protected def span {v' : ι' → E} (h : orthonormal 𝕜 v') (s : finset ι') : +protected def span [decidable_eq E] {v' : ι' → E} (h : orthonormal 𝕜 v') (s : finset ι') : orthonormal_basis s 𝕜 (span 𝕜 (s.image v' : set E)) := let e₀' : basis s 𝕜 _ := basis.span (h.linear_independent.comp (coe : s → ι') subtype.coe_injective), @@ -421,7 +422,8 @@ let in e₀.map φ.symm -@[simp] protected lemma span_apply {v' : ι' → E} (h : orthonormal 𝕜 v') (s : finset ι') (i : s) : +@[simp] protected lemma span_apply [decidable_eq E] + {v' : ι' → E} (h : orthonormal 𝕜 v') (s : finset ι') (i : s) : (orthonormal_basis.span h s i : E) = v' i := by simp only [orthonormal_basis.span, basis.span_apply, linear_isometry_equiv.of_eq_symm, orthonormal_basis.map_apply, orthonormal_basis.coe_mk, diff --git a/src/combinatorics/simple_graph/basic.lean b/src/combinatorics/simple_graph/basic.lean index 2434050c02421..dc4a24b882697 100644 --- a/src/combinatorics/simple_graph/basic.lean +++ b/src/combinatorics/simple_graph/basic.lean @@ -606,13 +606,16 @@ attribute [mono] edge_finset_mono edge_finset_strict_mono @[simp] lemma edge_finset_bot : (⊥ : simple_graph V).edge_finset = ∅ := by simp [edge_finset] -@[simp] lemma edge_finset_sup : (G₁ ⊔ G₂).edge_finset = G₁.edge_finset ∪ G₂.edge_finset := +@[simp] lemma edge_finset_sup [decidable_eq V] : + (G₁ ⊔ G₂).edge_finset = G₁.edge_finset ∪ G₂.edge_finset := by simp [edge_finset] -@[simp] lemma edge_finset_inf : (G₁ ⊓ G₂).edge_finset = G₁.edge_finset ∩ G₂.edge_finset := +@[simp] lemma edge_finset_inf [decidable_eq V] : + (G₁ ⊓ G₂).edge_finset = G₁.edge_finset ∩ G₂.edge_finset := by simp [edge_finset] -@[simp] lemma edge_finset_sdiff : (G₁ \ G₂).edge_finset = G₁.edge_finset \ G₂.edge_finset := +@[simp] lemma edge_finset_sdiff [decidable_eq V] : + (G₁ \ G₂).edge_finset = G₁.edge_finset \ G₂.edge_finset := by simp [edge_finset] lemma edge_finset_card : G.edge_finset.card = fintype.card G.edge_set := set.to_finset_card _ diff --git a/src/data/dfinsupp/basic.lean b/src/data/dfinsupp/basic.lean index e69a8c4f99efe..8d649239ef552 100644 --- a/src/data/dfinsupp/basic.lean +++ b/src/data/dfinsupp/basic.lean @@ -1221,7 +1221,8 @@ begin sigma_curry_apply, smul_apply] end -@[simp] lemma sigma_curry_single [Π i j, has_zero (δ i j)] (ij : Σ i, α i) (x : δ ij.1 ij.2) : +@[simp] lemma sigma_curry_single [decidable_eq ι] [Π i, decidable_eq (α i)] + [Π i j, has_zero (δ i j)] (ij : Σ i, α i) (x : δ ij.1 ij.2) : @sigma_curry _ _ _ _ (single ij x) = single ij.1 (single ij.2 x : Π₀ j, δ ij.1 j) := begin obtain ⟨i, j⟩ := ij, @@ -1240,7 +1241,8 @@ end /--The natural map between `Π₀ i (j : α i), δ i j` and `Π₀ (i : Σ i, α i), δ i.1 i.2`, inverse of `curry`.-/ -noncomputable def sigma_uncurry [Π i j, has_zero (δ i j)] (f : Π₀ i j, δ i j) : +def sigma_uncurry [Π i j, has_zero (δ i j)] + [Π i, decidable_eq (α i)] [Π i j (x : δ i j), decidable (x ≠ 0)] (f : Π₀ i j, δ i j) : Π₀ (i : Σ i, _), δ i.1 i.2 := { to_fun := λ i, f i.1 i.2, support' := f.support'.map $ λ s, @@ -1255,24 +1257,32 @@ noncomputable def sigma_uncurry [Π i j, has_zero (δ i j)] (f : Π₀ i j, δ i rw [hi, zero_apply] } end⟩ } -@[simp] lemma sigma_uncurry_apply [Π i j, has_zero (δ i j)] (f : Π₀ i j, δ i j) (i : ι) (j : α i) : +@[simp] lemma sigma_uncurry_apply [Π i j, has_zero (δ i j)] + [Π i, decidable_eq (α i)] [Π i j (x : δ i j), decidable (x ≠ 0)] + (f : Π₀ i j, δ i j) (i : ι) (j : α i) : sigma_uncurry f ⟨i, j⟩ = f i j := rfl -@[simp] lemma sigma_uncurry_zero [Π i j, has_zero (δ i j)] : +@[simp] lemma sigma_uncurry_zero [Π i j, has_zero (δ i j)] + [Π i, decidable_eq (α i)] [Π i j (x : δ i j), decidable (x ≠ 0)]: sigma_uncurry (0 : Π₀ i j, δ i j) = 0 := rfl -@[simp] lemma sigma_uncurry_add [Π i j, add_zero_class (δ i j)] (f g : Π₀ i j, δ i j) : +@[simp] lemma sigma_uncurry_add [Π i j, add_zero_class (δ i j)] + [Π i, decidable_eq (α i)] [Π i j (x : δ i j), decidable (x ≠ 0)] + (f g : Π₀ i j, δ i j) : sigma_uncurry (f + g) = sigma_uncurry f + sigma_uncurry g := coe_fn_injective rfl @[simp] lemma sigma_uncurry_smul [monoid γ] [Π i j, add_monoid (δ i j)] + [Π i, decidable_eq (α i)] [Π i j (x : δ i j), decidable (x ≠ 0)] [Π i j, distrib_mul_action γ (δ i j)] (r : γ) (f : Π₀ i j, δ i j) : sigma_uncurry (r • f) = r • sigma_uncurry f := coe_fn_injective rfl -@[simp] lemma sigma_uncurry_single [Π i j, has_zero (δ i j)] (i) (j : α i) (x : δ i j) : +@[simp] lemma sigma_uncurry_single [Π i j, has_zero (δ i j)] + [decidable_eq ι] [Π i, decidable_eq (α i)] [Π i j (x : δ i j), decidable (x ≠ 0)] + (i) (j : α i) (x : δ i j) : sigma_uncurry (single i (single j x : Π₀ (j : α i), δ i j)) = single ⟨i, j⟩ x:= begin ext ⟨i', j'⟩, @@ -1291,7 +1301,8 @@ end /--The natural bijection between `Π₀ (i : Σ i, α i), δ i.1 i.2` and `Π₀ i (j : α i), δ i j`. This is the dfinsupp version of `equiv.Pi_curry`. -/ -noncomputable def sigma_curry_equiv [Π i j, has_zero (δ i j)] : +noncomputable def sigma_curry_equiv [Π i j, has_zero (δ i j)] + [Π i, decidable_eq (α i)] [Π i j (x : δ i j), decidable (x ≠ 0)] : (Π₀ (i : Σ i, _), δ i.1 i.2) ≃ Π₀ i j, δ i j := { to_fun := sigma_curry, inv_fun := sigma_uncurry, diff --git a/src/data/finsupp/alist.lean b/src/data/finsupp/alist.lean index 48540fb78e393..9ede0597051b9 100644 --- a/src/data/finsupp/alist.lean +++ b/src/data/finsupp/alist.lean @@ -91,6 +91,7 @@ by { classical, simp [←alist.insert_empty] } @[simp] lemma _root_.finsupp.to_alist_lookup_finsupp (f : α →₀ M) : f.to_alist.lookup_finsupp = f := begin ext, + classical, by_cases h : f a = 0, { suffices : f.to_alist.lookup a = none, { simp [h, this] }, diff --git a/src/data/finsupp/basic.lean b/src/data/finsupp/basic.lean index f2cba85f09fa7..516f5255164bd 100644 --- a/src/data/finsupp/basic.lean +++ b/src/data/finsupp/basic.lean @@ -901,7 +901,7 @@ prod_bij (λp _, p.val) (λ _, by classical; exact mem_subtype.1) (λ _ _, rfl) (λ _ _ _ _, subtype.eq) - (λ b hb, ⟨⟨b, hp b hb⟩, mem_subtype.2 hb, rfl⟩) + (λ b hb, ⟨⟨b, hp b hb⟩, by classical; exact mem_subtype.2 hb, rfl⟩) end zero diff --git a/src/data/finsupp/defs.lean b/src/data/finsupp/defs.lean index 5ddfe84c184a1..4ba86d3d97737 100644 --- a/src/data/finsupp/defs.lean +++ b/src/data/finsupp/defs.lean @@ -222,21 +222,21 @@ def single (a : α) (b : M) : α →₀ M := end } lemma single_apply [decidable (a = a')] : single a b a' = if a = a' then b else 0 := -by { simp_rw [@eq_comm _ a a'], convert pi.single_apply _ _ _, } +by { classical, simp_rw [@eq_comm _ a a'], convert pi.single_apply _ _ _, } lemma single_apply_left {f : α → β} (hf : function.injective f) (x z : α) (y : M) : single (f x) y (f z) = single x y z := -by simp only [single_apply, hf.eq_iff] +by { classical, simp only [single_apply, hf.eq_iff] } lemma single_eq_indicator : ⇑(single a b) = set.indicator {a} (λ _, b) := -by { ext, simp [single_apply, set.indicator, @eq_comm _ a] } +by { classical, ext, simp [single_apply, set.indicator, @eq_comm _ a] } @[simp] lemma single_eq_same : (single a b : α →₀ M) a = b := -pi.single_eq_same a b +by { classical, exact pi.single_eq_same a b } @[simp] lemma single_eq_of_ne (h : a ≠ a') : (single a b : α →₀ M) a' = 0 := -pi.single_eq_of_ne' h _ +by { classical, exact pi.single_eq_of_ne' h _ } lemma single_eq_update [decidable_eq α] (a : α) (b : M) : ⇑(single a b) = function.update 0 a b := by rw [single_eq_indicator, ← set.piecewise_eq_indicator, set.piecewise_singleton] @@ -245,12 +245,15 @@ lemma single_eq_pi_single [decidable_eq α] (a : α) (b : M) : ⇑(single a b) = single_eq_update a b @[simp] lemma single_zero (a : α) : (single a 0 : α →₀ M) = 0 := -coe_fn_injective $ by simpa only [single_eq_update, coe_zero] - using function.update_eq_self a (0 : α → M) +coe_fn_injective $ begin + classical, + simpa only [single_eq_update, coe_zero] using function.update_eq_self a (0 : α → M) +end lemma single_of_single_apply (a a' : α) (b : M) : single a ((single a' b) a) = single a' (single a' b) a := begin + classical, rw [single_apply, single_apply], ext, split_ifs, @@ -259,10 +262,10 @@ begin end lemma support_single_ne_zero (a : α) (hb : b ≠ 0) : (single a b).support = {a} := -if_neg hb +by { classical, exact if_neg hb } lemma support_single_subset : (single a b).support ⊆ {a} := -show ite _ _ _ ⊆ _, by split_ifs; [exact empty_subset _, exact subset.refl _] +by { classical, show ite _ _ _ ⊆ _, split_ifs; [exact empty_subset _, exact subset.refl _] } lemma single_apply_mem (x) : single a b x ∈ ({0, b} : set M) := by rcases em (a = x) with (rfl|hx); [simp, simp [single_eq_of_ne hx]] @@ -334,7 +337,7 @@ by rw [support_single_ne_zero _ hb, support_single_ne_zero _ hb', disjoint_singl by simp [ext_iff, single_eq_indicator] lemma single_swap (a₁ a₂ : α) (b : M) : single a₁ b a₂ = single a₂ b a₁ := -by simp only [single_apply]; ac_refl +by { classical, simp only [single_apply], ac_refl } instance [nonempty α] [nontrivial M] : nontrivial (α →₀ M) := begin diff --git a/src/data/finsupp/to_dfinsupp.lean b/src/data/finsupp/to_dfinsupp.lean index 9254fc3fa56b0..e4b018ab42d15 100644 --- a/src/data/finsupp/to_dfinsupp.lean +++ b/src/data/finsupp/to_dfinsupp.lean @@ -257,6 +257,7 @@ begin suffices : finsupp.single (⟨i, a⟩ : Σ i, η i) n ⟨j, b⟩ = 0, { simp [split_apply, dif_neg h, this] }, have H : (⟨i, a⟩ : Σ i, η i) ≠ ⟨j, b⟩ := by simp [h], + classical, rw [finsupp.single_apply, if_neg H] end diff --git a/src/data/fintype/basic.lean b/src/data/fintype/basic.lean index 5603a0ecd5c40..94075464d09fc 100644 --- a/src/data/fintype/basic.lean +++ b/src/data/fintype/basic.lean @@ -756,7 +756,7 @@ noncomputable def finset_equiv_set [fintype α] : finset α ≃ set α := { to_fun := coe, inv_fun := by { classical, exact λ s, s.to_finset }, left_inv := λ s, by convert finset.to_finset_coe s, - right_inv := λ s, s.coe_to_finset } + right_inv := λ s, by { classical, exact s.coe_to_finset } } @[simp] lemma finset_equiv_set_apply [fintype α] (s : finset α) : finset_equiv_set s = s := rfl diff --git a/src/data/fintype/sum.lean b/src/data/fintype/sum.lean index 5359cdd8db246..fe14ab22471cc 100644 --- a/src/data/fintype/sum.lean +++ b/src/data/fintype/sum.lean @@ -39,7 +39,7 @@ def fintype_of_fintype_ne (a : α) (h : fintype {b // b ≠ a}) : fintype α := fintype.of_bijective (sum.elim (coe : {b // b = a} → α) (coe : {b // b ≠ a} → α)) $ by { classical, exact (equiv.sum_compl (= a)).bijective } -lemma image_subtype_ne_univ_eq_image_erase [fintype α] (k : β) (b : α → β) : +lemma image_subtype_ne_univ_eq_image_erase [fintype α] [decidable_eq β] (k : β) (b : α → β) : image (λ i : {a // b a ≠ k}, b ↑i) univ = (image b univ).erase k := begin apply subset_antisymm, @@ -53,7 +53,7 @@ begin exact ⟨⟨a, ne_of_mem_erase hi⟩, mem_univ _, rfl⟩ } end -lemma image_subtype_univ_ssubset_image_univ [fintype α] (k : β) (b : α → β) +lemma image_subtype_univ_ssubset_image_univ [fintype α] [decidable_eq β] (k : β) (b : α → β) (hk : k ∈ image b univ) (p : β → Prop) [decidable_pred p] (hp : ¬ p k) : image (λ i : {a // p (b a)}, b ↑i) univ ⊂ image b univ := begin @@ -72,7 +72,7 @@ end /-- Any injection from a finset `s` in a fintype `α` to a finset `t` of the same cardinality as `α` can be extended to a bijection between `α` and `t`. -/ -lemma finset.exists_equiv_extend_of_card_eq [fintype α] {t : finset β} +lemma finset.exists_equiv_extend_of_card_eq [fintype α] [decidable_eq β] {t : finset β} (hαt : fintype.card α = t.card) {s : finset α} {f : α → β} (hfst : s.image f ⊆ t) (hfs : set.inj_on f s) : ∃ g : α ≃ t, ∀ i ∈ s, (g i : β) = f i := diff --git a/src/data/pi/lex.lean b/src/data/pi/lex.lean index 6e7760e4c1e5d..a6f2d78919bc6 100644 --- a/src/data/pi/lex.lean +++ b/src/data/pi/lex.lean @@ -173,8 +173,11 @@ end⟩ lemma lex.no_max_order' [preorder ι] [Π i, has_lt (β i)] (i : ι) [no_max_order (β i)] : no_max_order (lex (Π i, β i)) := -⟨λ a, let ⟨b, hb⟩ := exists_gt (a i) in ⟨a.update i b, i, - λ j hj, (a.update_noteq hj.ne b).symm, by rwa a.update_same i b⟩⟩ +⟨λ a, begin + classical, + obtain ⟨b, hb⟩ := exists_gt (a i), + exact ⟨a.update i b, i, λ j hj, (a.update_noteq hj.ne b).symm, by rwa a.update_same i b⟩ +end⟩ instance [linear_order ι] [is_well_order ι (<)] [nonempty ι] [Π i, partial_order (β i)] [Π i, no_max_order (β i)] : diff --git a/src/group_theory/divisible.lean b/src/group_theory/divisible.lean index 63185f2f54c8f..d4678db9e4375 100644 --- a/src/group_theory/divisible.lean +++ b/src/group_theory/divisible.lean @@ -117,7 +117,7 @@ noncomputable def rootable_by_of_pow_left_surj rootable_by A α := { root := λ a n, @dite _ (n = 0) (classical.dec _) (λ _, (1 : A)) (λ hn, (H hn a).some), root_zero := λ _, by classical; exact dif_pos rfl, - root_cancel := λ n a hn, by rw dif_neg hn; exact (H hn a).some_spec } + root_cancel := λ n a hn, by { classical, rw dif_neg hn, exact (H hn a).some_spec } } section pi diff --git a/src/linear_algebra/free_module/ideal_quotient.lean b/src/linear_algebra/free_module/ideal_quotient.lean index feb64f957bec7..ab6d4937ceef9 100644 --- a/src/linear_algebra/free_module/ideal_quotient.lean +++ b/src/linear_algebra/free_module/ideal_quotient.lean @@ -96,5 +96,5 @@ let b := module.free.choose_basis ℤ S, a := I.smith_coeffs b hI, e := I.quotient_equiv_pi_zmod b hI in by haveI : ∀ i, ne_zero (a i).nat_abs := - (λ i, ⟨int.nat_abs_ne_zero_of_ne_zero (ideal.smith_coeffs_ne_zero b I hI i)⟩); + (λ i, ⟨int.nat_abs_ne_zero_of_ne_zero (ideal.smith_coeffs_ne_zero b I hI i)⟩); classical; exact fintype.of_equiv (Π i, zmod (a i).nat_abs) e.symm diff --git a/src/number_theory/legendre_symbol/mul_character.lean b/src/number_theory/legendre_symbol/mul_character.lean index 3c7903e9736f3..648797dd56d15 100644 --- a/src/number_theory/legendre_symbol/mul_character.lean +++ b/src/number_theory/legendre_symbol/mul_character.lean @@ -104,6 +104,7 @@ def trivial : mul_char R R' := map_nonunit' := by { intros a ha, simp only [ha, if_false], }, map_one' := by simp only [is_unit_one, if_true], map_mul' := by { intros x y, + classical, simp only [is_unit.mul_iff, boole_mul], split_ifs; tauto, } } @@ -175,6 +176,7 @@ def of_unit_hom (f : Rˣ →* R'ˣ) : mul_char R R' := simp only [h1, dif_pos, units.coe_eq_one, map_one, is_unit_one], }, map_mul' := begin + classical, intros x y, by_cases hx : is_unit x, { simp only [hx, is_unit.mul_iff, true_and, dif_pos], @@ -249,7 +251,7 @@ instance inhabited : inhabited (mul_char R R') := ⟨1⟩ /-- Evaluation of the trivial character -/ @[simp] lemma one_apply_coe (a : Rˣ) : (1 : mul_char R R') a = 1 := -dif_pos a.is_unit +by { classical, exact dif_pos a.is_unit } /-- Multiplication of multiplicative characters. (This needs the target to be commutative.) -/ def mul (χ χ' : mul_char R R') : mul_char R R' := diff --git a/src/order/interval.lean b/src/order/interval.lean index 23dd0ca388220..ece0f3aee0dea 100644 --- a/src/order/interval.lean +++ b/src/order/interval.lean @@ -446,7 +446,8 @@ by classical; exact { Sup := λ S, if h : S ⊆ {⊥} then ⊥ else some end, ..interval.lattice, ..interval.bounded_order } -@[simp, norm_cast] lemma coe_Inf (S : set (interval α)) : ↑(Inf S) = ⋂ s ∈ S, (s : set α) := +@[simp, norm_cast] lemma coe_Inf [@decidable_rel α (≤)] (S : set (interval α)) : + ↑(Inf S) = ⋂ s ∈ S, (s : set α) := begin change coe (dite _ _ _) = _, split_ifs, @@ -463,10 +464,11 @@ begin exact h (λ s ha t hb, (hx _ ha).1.trans (hx _ hb).2) } end -@[simp, norm_cast] lemma coe_infi (f : ι → interval α) : ↑(⨅ i, f i) = ⋂ i, (f i : set α) := +@[simp, norm_cast] lemma coe_infi [@decidable_rel α (≤)] (f : ι → interval α) : + ↑(⨅ i, f i) = ⋂ i, (f i : set α) := by simp [infi] -@[simp, norm_cast] lemma coe_infi₂ (f : Π i, κ i → interval α) : +@[simp, norm_cast] lemma coe_infi₂ [@decidable_rel α (≤)] (f : Π i, κ i → interval α) : ↑(⨅ i j, f i j) = ⋂ i j, (f i j : set α) := by simp_rw [coe_infi] diff --git a/src/order/max.lean b/src/order/max.lean index ad59fa0fa0b6f..01aa6e0e063d5 100644 --- a/src/order/max.lean +++ b/src/order/max.lean @@ -95,6 +95,7 @@ end⟩ instance [nonempty ι] [Π i, preorder (π i)] [Π i, no_min_order (π i)] : no_min_order (Π i, π i) := ⟨λ a, begin + classical, obtain ⟨b, hb⟩ := exists_lt (a $ classical.arbitrary _), exact ⟨_, update_lt_self_iff.2 hb⟩, end⟩ diff --git a/src/order/order_iso_nat.lean b/src/order/order_iso_nat.lean index e7eec2a4798cc..505d25a4e92ee 100644 --- a/src/order/order_iso_nat.lean +++ b/src/order/order_iso_nat.lean @@ -107,17 +107,21 @@ by { classical, exact rel_iso.of_surjective (rel_embedding.order_embedding_of_lt variable {s} @[simp] -lemma coe_order_embedding_of_set : ⇑(order_embedding_of_set s) = coe ∘ subtype.of_nat s := rfl +lemma coe_order_embedding_of_set [decidable_pred (∈ s)] : + ⇑(order_embedding_of_set s) = coe ∘ subtype.of_nat s := rfl -lemma order_embedding_of_set_apply {n : ℕ} : order_embedding_of_set s n = subtype.of_nat s n := rfl +lemma order_embedding_of_set_apply [decidable_pred (∈ s)] {n : ℕ} : + order_embedding_of_set s n = subtype.of_nat s n := rfl @[simp] -lemma subtype.order_iso_of_nat_apply {n : ℕ} : subtype.order_iso_of_nat s n = subtype.of_nat s n := -by simp [subtype.order_iso_of_nat] +lemma subtype.order_iso_of_nat_apply [decidable_pred (∈ s)] {n : ℕ} : + subtype.order_iso_of_nat s n = subtype.of_nat s n := +by { simp [subtype.order_iso_of_nat], congr } variable (s) -lemma order_embedding_of_set_range : set.range (nat.order_embedding_of_set s) = s := +lemma order_embedding_of_set_range [decidable_pred (∈ s)] : + set.range (nat.order_embedding_of_set s) = s := subtype.coe_comp_of_nat_range theorem exists_subseq_of_forall_mem_union {s t : set α} (e : ℕ → α) (he : ∀ n, e n ∈ s ∪ t) : diff --git a/src/order/prop_instances.lean b/src/order/prop_instances.lean index e27289e1d8153..bce4033a2e8e6 100644 --- a/src/order/prop_instances.lean +++ b/src/order/prop_instances.lean @@ -59,6 +59,7 @@ lemma disjoint_iff [Π i, order_bot (α' i)] {f g : Π i, α' i} : begin split, { intros h i x hf hg, + classical, refine (update_le_iff.mp $ -- this line doesn't work h (update_le_iff.mpr ⟨hf, λ _ _, _⟩) (update_le_iff.mpr ⟨hg, λ _ _, _⟩)).1, @@ -85,5 +86,6 @@ codisjoint_iff_le_sup.trans $ forall_const _ @[simp] lemma Prop.is_compl_iff {P Q : Prop} : is_compl P Q ↔ ¬(P ↔ Q) := begin rw [is_compl_iff, Prop.disjoint_iff, Prop.codisjoint_iff, not_iff], + classical, tauto, end diff --git a/src/ring_theory/bezout.lean b/src/ring_theory/bezout.lean index 7b2c4d6f9fdc1..f2e588af810d1 100644 --- a/src/ring_theory/bezout.lean +++ b/src/ring_theory/bezout.lean @@ -95,7 +95,8 @@ local attribute [instance] to_gcd_domain -- Note that the proof, despite being `infer_instance`, depends on the `local attribute [instance]` -- lemma above, and is thus necessary to be restated. @[priority 100] -instance [is_domain R] [is_bezout R] : is_integrally_closed R := infer_instance +instance [is_domain R] [is_bezout R] : is_integrally_closed R := +by classical; exact gcd_monoid.to_is_integrally_closed lemma _root_.function.surjective.is_bezout {S : Type v} [comm_ring S] (f : R →+* S) (hf : function.surjective f) [is_bezout R] : is_bezout S := diff --git a/src/ring_theory/discriminant.lean b/src/ring_theory/discriminant.lean index 1d851a10e5461..1d156e5f4cf03 100644 --- a/src/ring_theory/discriminant.lean +++ b/src/ring_theory/discriminant.lean @@ -127,6 +127,7 @@ begin { simp [discr] }, { have := span_eq_top_of_linear_independent_of_card_eq_finrank b.linear_independent (finrank_eq_card_basis b).symm, + classical, rw [discr_def, trace_matrix_def], simp_rw [← basis.mk_apply b.linear_independent this.ge], rw [← trace_matrix_def, trace_matrix_of_basis, ← bilin_form.nondegenerate_iff_det_ne_zero], diff --git a/src/set_theory/game/basic.lean b/src/set_theory/game/basic.lean index 5becb01f79721..c440ab66e49f4 100644 --- a/src/set_theory/game/basic.lean +++ b/src/set_theory/game/basic.lean @@ -590,16 +590,17 @@ noncomputable instance : has_inv pgame := noncomputable instance : has_div pgame := ⟨λ x y, x * y⁻¹⟩ -theorem inv_eq_of_equiv_zero {x : pgame} (h : x ≈ 0) : x⁻¹ = 0 := if_pos h +theorem inv_eq_of_equiv_zero {x : pgame} (h : x ≈ 0) : x⁻¹ = 0 := +by { classical, exact if_pos h } @[simp] theorem inv_zero : (0 : pgame)⁻¹ = 0 := inv_eq_of_equiv_zero (equiv_refl _) theorem inv_eq_of_pos {x : pgame} (h : 0 < x) : x⁻¹ = inv' x := -(if_neg h.lf.not_equiv').trans (if_pos h) +by { classical, exact (if_neg h.lf.not_equiv').trans (if_pos h) } theorem inv_eq_of_lf_zero {x : pgame} (h : x ⧏ 0) : x⁻¹ = -inv' (-x) := -(if_neg h.not_equiv).trans (if_neg h.not_gt) +by { classical, exact (if_neg h.not_equiv).trans (if_neg h.not_gt) } /-- `1⁻¹` has exactly the same moves as `1`. -/ def inv_one : 1⁻¹ ≡r 1 := From 2987594a0b88ccb096b11b0f4b17923a9e11df02 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Wed, 25 Jan 2023 06:25:44 +0000 Subject: [PATCH 0345/1291] refactor(topology): move `pcontinuous` etc to a new file (#18288) Move `pcontinuous` and lemmas about `ptendsto` to a new file. There are some issues with porting `data.pfun` to Lean 4, so this PR makes it possible to port topology without waiting for `pfun`. --- src/combinatorics/hall/basic.lean | 1 + src/topology/basic.lean | 49 --------------------- src/topology/continuous_on.lean | 4 -- src/topology/partial.lean | 73 +++++++++++++++++++++++++++++++ 4 files changed, 74 insertions(+), 53 deletions(-) create mode 100644 src/topology/partial.lean diff --git a/src/combinatorics/hall/basic.lean b/src/combinatorics/hall/basic.lean index deabb27394632..2251daea1479e 100644 --- a/src/combinatorics/hall/basic.lean +++ b/src/combinatorics/hall/basic.lean @@ -5,6 +5,7 @@ Authors: Alena Gusakov, Bhavik Mehta, Kyle Miller -/ import combinatorics.hall.finite import topology.category.Top.limits +import data.rel /-! # Hall's Marriage Theorem diff --git a/src/topology/basic.lean b/src/topology/basic.lean index e3784ae39ea07..9b62d1a97598b 100644 --- a/src/topology/basic.lean +++ b/src/topology/basic.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro, Jeremy Avigad -/ import order.filter.ultrafilter -import order.filter.partial import algebra.support import order.filter.lift @@ -817,22 +816,6 @@ theorem all_mem_nhds_filter (x : α) (f : set α → set β) (hf : ∀ s t, s (∀ s ∈ 𝓝 x, f s ∈ l) ↔ (∀ s, is_open s → x ∈ s → f s ∈ l) := all_mem_nhds _ _ (λ s t ssubt h, mem_of_superset h (hf s t ssubt)) -theorem rtendsto_nhds {r : rel β α} {l : filter β} {a : α} : - rtendsto r l (𝓝 a) ↔ (∀ s, is_open s → a ∈ s → r.core s ∈ l) := -all_mem_nhds_filter _ _ (λ s t, id) _ - -theorem rtendsto'_nhds {r : rel β α} {l : filter β} {a : α} : - rtendsto' r l (𝓝 a) ↔ (∀ s, is_open s → a ∈ s → r.preimage s ∈ l) := -by { rw [rtendsto'_def], apply all_mem_nhds_filter, apply rel.preimage_mono } - -theorem ptendsto_nhds {f : β →. α} {l : filter β} {a : α} : - ptendsto f l (𝓝 a) ↔ (∀ s, is_open s → a ∈ s → f.core s ∈ l) := -rtendsto_nhds - -theorem ptendsto'_nhds {f : β →. α} {l : filter β} {a : α} : - ptendsto' f l (𝓝 a) ↔ (∀ s, is_open s → a ∈ s → f.preimage s ∈ l) := -rtendsto'_nhds - theorem tendsto_nhds {f : β → α} {l : filter β} {a : α} : tendsto f l (𝓝 a) ↔ (∀ s, is_open s → a ∈ s → f ⁻¹' s ∈ l) := all_mem_nhds_filter _ _ (λ s t h, preimage_mono h) _ @@ -1431,38 +1414,6 @@ lemma continuous.frontier_preimage_subset frontier (f ⁻¹' t) ⊆ f ⁻¹' (frontier t) := diff_subset_diff (hf.closure_preimage_subset t) (preimage_interior_subset_interior_preimage hf) -/-! ### Continuity and partial functions -/ - -/-- Continuity of a partial function -/ -def pcontinuous (f : α →. β) := ∀ s, is_open s → is_open (f.preimage s) - -lemma open_dom_of_pcontinuous {f : α →. β} (h : pcontinuous f) : is_open f.dom := -by rw [←pfun.preimage_univ]; exact h _ is_open_univ - -lemma pcontinuous_iff' {f : α →. β} : - pcontinuous f ↔ ∀ {x y} (h : y ∈ f x), ptendsto' f (𝓝 x) (𝓝 y) := -begin - split, - { intros h x y h', - simp only [ptendsto'_def, mem_nhds_iff], - rintros s ⟨t, tsubs, opent, yt⟩, - exact ⟨f.preimage t, pfun.preimage_mono _ tsubs, h _ opent, ⟨y, yt, h'⟩⟩ }, - intros hf s os, - rw is_open_iff_nhds, - rintros x ⟨y, ys, fxy⟩ t, - rw [mem_principal], - assume h : f.preimage s ⊆ t, - change t ∈ 𝓝 x, - apply mem_of_superset _ h, - have h' : ∀ s ∈ 𝓝 y, f.preimage s ∈ 𝓝 x, - { intros s hs, - have : ptendsto' f (𝓝 x) (𝓝 y) := hf fxy, - rw ptendsto'_def at this, - exact this s hs }, - show f.preimage s ∈ 𝓝 x, - apply h', rw mem_nhds_iff, exact ⟨s, set.subset.refl _, os, ys⟩ -end - /-- If a continuous map `f` maps `s` to `t`, then it maps `closure s` to `closure t`. -/ lemma set.maps_to.closure {s : set α} {t : set β} {f : α → β} (h : maps_to f s t) (hc : continuous f) : maps_to f (closure s) (closure t) := diff --git a/src/topology/continuous_on.lean b/src/topology/continuous_on.lean index 8af2282f3e25d..35dab90838e09 100644 --- a/src/topology/continuous_on.lean +++ b/src/topology/continuous_on.lean @@ -615,10 +615,6 @@ ctsf.tendsto_nhds_within_image.le_comap comap f (𝓝[range f] y) = comap f (𝓝 y) := comap_inf_principal_range -theorem continuous_within_at_iff_ptendsto_res (f : α → β) {x : α} {s : set α} : - continuous_within_at f s x ↔ ptendsto (pfun.res f s) (𝓝 x) (𝓝 (f x)) := -tendsto_iff_ptendsto _ _ _ _ - lemma continuous_iff_continuous_on_univ {f : α → β} : continuous f ↔ continuous_on f univ := by simp [continuous_iff_continuous_at, continuous_on, continuous_at, continuous_within_at, nhds_within_univ] diff --git a/src/topology/partial.lean b/src/topology/partial.lean new file mode 100644 index 0000000000000..c09f9b41ab63f --- /dev/null +++ b/src/topology/partial.lean @@ -0,0 +1,73 @@ +/- +Copyright (c) 2018 Jeremy Avigad. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jeremy Avigad +-/ +import topology.continuous_on +import order.filter.partial + +/-! +# Partial functions and topological spaces + +In this file we prove properties of `filter.ptendsto` etc in topological spaces. We also introduce +`pcontinuous`, a version of `continuous` for partially defined functions. +-/ + +open filter +open_locale topological_space + +variables {α β : Type*} [topological_space α] + +theorem rtendsto_nhds {r : rel β α} {l : filter β} {a : α} : + rtendsto r l (𝓝 a) ↔ (∀ s, is_open s → a ∈ s → r.core s ∈ l) := +all_mem_nhds_filter _ _ (λ s t, id) _ + +theorem rtendsto'_nhds {r : rel β α} {l : filter β} {a : α} : + rtendsto' r l (𝓝 a) ↔ (∀ s, is_open s → a ∈ s → r.preimage s ∈ l) := +by { rw [rtendsto'_def], apply all_mem_nhds_filter, apply rel.preimage_mono } + +theorem ptendsto_nhds {f : β →. α} {l : filter β} {a : α} : + ptendsto f l (𝓝 a) ↔ (∀ s, is_open s → a ∈ s → f.core s ∈ l) := +rtendsto_nhds + +theorem ptendsto'_nhds {f : β →. α} {l : filter β} {a : α} : + ptendsto' f l (𝓝 a) ↔ (∀ s, is_open s → a ∈ s → f.preimage s ∈ l) := +rtendsto'_nhds + +/-! ### Continuity and partial functions -/ + +variable [topological_space β] + +/-- Continuity of a partial function -/ +def pcontinuous (f : α →. β) := ∀ s, is_open s → is_open (f.preimage s) + +lemma open_dom_of_pcontinuous {f : α →. β} (h : pcontinuous f) : is_open f.dom := +by rw [←pfun.preimage_univ]; exact h _ is_open_univ + +lemma pcontinuous_iff' {f : α →. β} : + pcontinuous f ↔ ∀ {x y} (h : y ∈ f x), ptendsto' f (𝓝 x) (𝓝 y) := +begin + split, + { intros h x y h', + simp only [ptendsto'_def, mem_nhds_iff], + rintros s ⟨t, tsubs, opent, yt⟩, + exact ⟨f.preimage t, pfun.preimage_mono _ tsubs, h _ opent, ⟨y, yt, h'⟩⟩ }, + intros hf s os, + rw is_open_iff_nhds, + rintros x ⟨y, ys, fxy⟩ t, + rw [mem_principal], + assume h : f.preimage s ⊆ t, + change t ∈ 𝓝 x, + apply mem_of_superset _ h, + have h' : ∀ s ∈ 𝓝 y, f.preimage s ∈ 𝓝 x, + { intros s hs, + have : ptendsto' f (𝓝 x) (𝓝 y) := hf fxy, + rw ptendsto'_def at this, + exact this s hs }, + show f.preimage s ∈ 𝓝 x, + apply h', rw mem_nhds_iff, exact ⟨s, set.subset.refl _, os, ys⟩ +end + +theorem continuous_within_at_iff_ptendsto_res (f : α → β) {x : α} {s : set α} : + continuous_within_at f s x ↔ ptendsto (pfun.res f s) (𝓝 x) (𝓝 (f x)) := +tendsto_iff_ptendsto _ _ _ _ From 4d392a6c9c4539cbeca399b3ee0afea398fbd2eb Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Wed, 25 Jan 2023 08:27:41 +0000 Subject: [PATCH 0346/1291] chore(*): add mathlib4 synchronization comments (#18283) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.ring.fin` * `control.functor.multivariate` * `data.fin.tuple.sort` * `data.lazy_list.basic` * `data.set.countable` * `data.set.equitable` * `data.set.intervals.infinite` * `data.set.intervals.instances` * `order.filter.archimedean` * `order.filter.at_top_bot` * `order.filter.bases` * `order.filter.cofinite` * `order.filter.countable_Inter` * `order.filter.indicator_function` * `order.filter.lift` * `order.filter.n_ary` * `order.filter.pi` * `order.filter.small_sets` * `order.filter.ultrafilter` * `order.order_iso_nat` --- src/algebra/ring/fin.lean | 3 +++ src/control/functor/multivariate.lean | 3 +++ src/data/fin/tuple/sort.lean | 3 +++ src/data/lazy_list/basic.lean | 3 +++ src/data/set/countable.lean | 3 +++ src/data/set/equitable.lean | 3 +++ src/data/set/intervals/infinite.lean | 3 +++ src/data/set/intervals/instances.lean | 3 +++ src/order/filter/archimedean.lean | 3 +++ src/order/filter/at_top_bot.lean | 3 +++ src/order/filter/bases.lean | 3 +++ src/order/filter/cofinite.lean | 3 +++ src/order/filter/countable_Inter.lean | 3 +++ src/order/filter/indicator_function.lean | 3 +++ src/order/filter/lift.lean | 3 +++ src/order/filter/n_ary.lean | 3 +++ src/order/filter/pi.lean | 3 +++ src/order/filter/small_sets.lean | 3 +++ src/order/filter/ultrafilter.lean | 3 +++ src/order/order_iso_nat.lean | 3 +++ 20 files changed, 60 insertions(+) diff --git a/src/algebra/ring/fin.lean b/src/algebra/ring/fin.lean index c465c67f7a78e..0b9379c02c61c 100644 --- a/src/algebra/ring/fin.lean +++ b/src/algebra/ring/fin.lean @@ -10,6 +10,9 @@ import algebra.group.prod /-! # Rings and `fin` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file collects some basic results involving rings and the `fin` type ## Main results diff --git a/src/control/functor/multivariate.lean b/src/control/functor/multivariate.lean index e9199e5bc22b5..40e62aab0df6d 100644 --- a/src/control/functor/multivariate.lean +++ b/src/control/functor/multivariate.lean @@ -7,6 +7,9 @@ import data.fin.fin2 import data.typevec /-! +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Functors between the category of tuples of types, and the category Type diff --git a/src/data/fin/tuple/sort.lean b/src/data/fin/tuple/sort.lean index 38430667b8f63..27d0bcc70bfbf 100644 --- a/src/data/fin/tuple/sort.lean +++ b/src/data/fin/tuple/sort.lean @@ -12,6 +12,9 @@ import group_theory.perm.basic # Sorting tuples by their values +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given an `n`-tuple `f : fin n → α` where `α` is ordered, we may want to turn it into a sorted `n`-tuple. This file provides an API for doing so, with the sorted `n`-tuple given by diff --git a/src/data/lazy_list/basic.lean b/src/data/lazy_list/basic.lean index 2ffe03b984210..4b6762f43bc70 100644 --- a/src/data/lazy_list/basic.lean +++ b/src/data/lazy_list/basic.lean @@ -10,6 +10,9 @@ import data.lazy_list /-! ## Definitions on lazy lists +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains various definitions and proofs on lazy lists. TODO: move the `lazy_list.lean` file from core to mathlib. diff --git a/src/data/set/countable.lean b/src/data/set/countable.lean index 3882fc93f9cfe..87b99ae19018d 100644 --- a/src/data/set/countable.lean +++ b/src/data/set/countable.lean @@ -9,6 +9,9 @@ import logic.equiv.list /-! # Countable sets + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ noncomputable theory diff --git a/src/data/set/equitable.lean b/src/data/set/equitable.lean index e0687fb4ba4b5..4d3d766714651 100644 --- a/src/data/set/equitable.lean +++ b/src/data/set/equitable.lean @@ -9,6 +9,9 @@ import data.nat.basic /-! # Equitable functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines equitable functions. A function `f` is equitable on a set `s` if `f a₁ ≤ f a₂ + 1` for all `a₁, a₂ ∈ s`. This is mostly diff --git a/src/data/set/intervals/infinite.lean b/src/data/set/intervals/infinite.lean index d86ec7f436560..2851bcffdf423 100644 --- a/src/data/set/intervals/infinite.lean +++ b/src/data/set/intervals/infinite.lean @@ -8,6 +8,9 @@ import data.set.finite /-! # Infinitude of intervals +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Bounded intervals in dense orders are infinite, as are unbounded intervals in orders that are unbounded on the appropriate side. We also prove that an unbounded preorder is an infinite type. diff --git a/src/data/set/intervals/instances.lean b/src/data/set/intervals/instances.lean index 21a3ed791b359..b70ae02eee13e 100644 --- a/src/data/set/intervals/instances.lean +++ b/src/data/set/intervals/instances.lean @@ -9,6 +9,9 @@ import algebra.ring.regular /-! # Algebraic instances for unit intervals +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + For suitably structured underlying type `α`, we exhibit the structure of the unit intervals (`set.Icc`, `set.Ioc`, `set.Ioc`, and `set.Ioo`) from `0` to `1`. diff --git a/src/order/filter/archimedean.lean b/src/order/filter/archimedean.lean index 1734c6a0a5833..81fa27e031c80 100644 --- a/src/order/filter/archimedean.lean +++ b/src/order/filter/archimedean.lean @@ -9,6 +9,9 @@ import order.filter.at_top_bot /-! # `at_top` filter and archimedean (semi)rings/fields +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove that for a linear ordered archimedean semiring `R` and a function `f : α → ℕ`, the function `coe ∘ f : α → R` tends to `at_top` along a filter `l` if and only if so does `f`. We also prove that `coe : ℕ → R` tends to `at_top` along `at_top`, as well as version of these diff --git a/src/order/filter/at_top_bot.lean b/src/order/filter/at_top_bot.lean index e0cc32ed65b4b..086cf2580565c 100644 --- a/src/order/filter/at_top_bot.lean +++ b/src/order/filter/at_top_bot.lean @@ -13,6 +13,9 @@ import algebra.order.group.min_max /-! # `at_top` and `at_bot` filters on preorded sets, monoids and groups. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define the filters * `at_top`: corresponds to `n → +∞`; diff --git a/src/order/filter/bases.lean b/src/order/filter/bases.lean index 9e2ead76bddd3..2b8d36b1ce91d 100644 --- a/src/order/filter/bases.lean +++ b/src/order/filter/bases.lean @@ -10,6 +10,9 @@ import order.filter.prod /-! # Filter bases +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A filter basis `B : filter_basis α` on a type `α` is a nonempty collection of sets of `α` such that the intersection of two elements of this collection contains some element of the collection. Compared to filters, filter bases do not require that any set containing diff --git a/src/order/filter/cofinite.lean b/src/order/filter/cofinite.lean index a2f292575c2f5..a9257107c807b 100644 --- a/src/order/filter/cofinite.lean +++ b/src/order/filter/cofinite.lean @@ -9,6 +9,9 @@ import order.filter.pi /-! # The cofinite filter +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define `cofinite`: the filter of sets with finite complement diff --git a/src/order/filter/countable_Inter.lean b/src/order/filter/countable_Inter.lean index 29aa9e318bbe1..7c1b523d91fa1 100644 --- a/src/order/filter/countable_Inter.lean +++ b/src/order/filter/countable_Inter.lean @@ -9,6 +9,9 @@ import data.set.countable /-! # Filters with countable intersection property +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define `countable_Inter_filter` to be the class of filters with the following property: for any countable collection of sets `s ∈ l` their intersection belongs to `l` as well. diff --git a/src/order/filter/indicator_function.lean b/src/order/filter/indicator_function.lean index 62c22c87cbeca..31787168fdfd0 100644 --- a/src/order/filter/indicator_function.lean +++ b/src/order/filter/indicator_function.lean @@ -9,6 +9,9 @@ import order.filter.at_top_bot /-! # Indicator function and filters +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Properties of indicator functions involving `=ᶠ` and `≤ᶠ`. ## Tags diff --git a/src/order/filter/lift.lean b/src/order/filter/lift.lean index 4d53a096d8680..c8d7f8a1e8605 100644 --- a/src/order/filter/lift.lean +++ b/src/order/filter/lift.lean @@ -7,6 +7,9 @@ import order.filter.bases /-! # Lift filters along filter and set functions + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ open set diff --git a/src/order/filter/n_ary.lean b/src/order/filter/n_ary.lean index 70744ed0bacce..9e67f2e016d1b 100644 --- a/src/order/filter/n_ary.lean +++ b/src/order/filter/n_ary.lean @@ -8,6 +8,9 @@ import order.filter.prod /-! # N-ary maps of filter +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the binary and ternary maps of filters. This is mostly useful to define pointwise operations on filters. diff --git a/src/order/filter/pi.lean b/src/order/filter/pi.lean index ee86f301ba4ef..d069ed0dcd512 100644 --- a/src/order/filter/pi.lean +++ b/src/order/filter/pi.lean @@ -8,6 +8,9 @@ import order.filter.bases /-! # (Co)product of a family of filters +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define two filters on `Π i, α i` and prove some basic properties of these filters. * `filter.pi (f : Π i, filter (α i))` to be the maximal filter on `Π i, α i` such that diff --git a/src/order/filter/small_sets.lean b/src/order/filter/small_sets.lean index ddb4e5fd9a6a9..2abc032ef6e0d 100644 --- a/src/order/filter/small_sets.lean +++ b/src/order/filter/small_sets.lean @@ -9,6 +9,9 @@ import order.filter.at_top_bot /-! # The filter of small sets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the filter of small sets w.r.t. a filter `f`, which is the largest filter containing all powersets of members of `f`. diff --git a/src/order/filter/ultrafilter.lean b/src/order/filter/ultrafilter.lean index fba9849333928..0d528f34e26f7 100644 --- a/src/order/filter/ultrafilter.lean +++ b/src/order/filter/ultrafilter.lean @@ -9,6 +9,9 @@ import order.zorn_atoms /-! # Ultrafilters +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + An ultrafilter is a minimal (maximal in the set order) proper filter. In this file we define diff --git a/src/order/order_iso_nat.lean b/src/order/order_iso_nat.lean index 505d25a4e92ee..85145f7acda7c 100644 --- a/src/order/order_iso_nat.lean +++ b/src/order/order_iso_nat.lean @@ -12,6 +12,9 @@ import tactic.congrm /-! # Relation embeddings from the naturals +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file allows translation from monotone functions `ℕ → α` to order embeddings `ℕ ↪ α` and defines the limit value of an eventually-constant sequence. From f974ae84dfc9fea6a036a8f30f09414254e3bc40 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Wed, 25 Jan 2023 11:06:51 +0000 Subject: [PATCH 0347/1291] chore(*): golf (#18111) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit ### Sets * Add `set.maps_to_prod_map_diagonal`, `set.diagonal_nonempty`, and `set.diagonal_subset_iff`. ### Filters * Generalize and rename `nhds_eq_comap_uniformity_aux` to `filter.mem_comap_prod_mk`. * Add `set.nonempty.principal_ne_bot` and `filter.comap_id'`. * Rename `filter.has_basis.comp_of_surjective` to `filter.has_basis.comp_surjective`. ### Uniform spaces * Rename `monotone_comp_rel` to `monotone.comp_rel` to enable dot notation. * Add `nhds_eq_comap_uniformity'`. * Use `𝓝ˢ (diagonal γ)` instead of `⨆ x, 𝓝 (x, x)` in `uniform_space_of_compact_t2`. * Golf here and there. ### Mathlib 4 port Relevant parts are forward-ported in leanprover-community/mathlib4#1438 --- src/data/set/function.lean | 3 + src/data/set/prod.lean | 9 +- src/order/filter/bases.lean | 4 +- src/order/filter/basic.lean | 9 ++ src/topology/algebra/uniform_group.lean | 2 +- src/topology/uniform_space/basic.lean | 104 ++++++------------ src/topology/uniform_space/compact.lean | 69 ++++-------- src/topology/uniform_space/completion.lean | 4 +- src/topology/uniform_space/separation.lean | 4 +- .../uniform_space/uniform_embedding.lean | 2 +- 10 files changed, 80 insertions(+), 130 deletions(-) diff --git a/src/data/set/function.lean b/src/data/set/function.lean index ad181e2038170..f2bc48dc31da1 100644 --- a/src/data/set/function.lean +++ b/src/data/set/function.lean @@ -295,6 +295,9 @@ lemma maps_to_iff_exists_map_subtype : maps_to f s t ↔ ∃ g : s → t, ∀ x theorem maps_to' : maps_to f s t ↔ f '' s ⊆ t := image_subset_iff.symm +theorem maps_to_prod_map_diagonal : maps_to (prod.map f f) (diagonal α) (diagonal β) := +diagonal_subset_iff.2 $ λ x, rfl + lemma maps_to.subset_preimage {f : α → β} {s : set α} {t : set β} (hf : maps_to f s t) : s ⊆ f ⁻¹' t := hf diff --git a/src/data/set/prod.lean b/src/data/set/prod.lean index 1a5ec550b1712..03c8cff1d27db 100644 --- a/src/data/set/prod.lean +++ b/src/data/set/prod.lean @@ -331,6 +331,9 @@ lemma mem_diagonal (x : α) : (x, x) ∈ diagonal α := by simp [diagonal] @[simp] lemma mem_diagonal_iff {x : α × α} : x ∈ diagonal α ↔ x.1 = x.2 := iff.rfl +lemma diagonal_nonempty [nonempty α] : (diagonal α).nonempty := +nonempty.elim ‹_› $ λ x, ⟨_, mem_diagonal x⟩ + instance decidable_mem_diagonal [h : decidable_eq α] (x : α × α) : decidable (x ∈ diagonal α) := h x.1 x.2 @@ -340,9 +343,11 @@ by { ext ⟨⟨x, hx⟩, ⟨y, hy⟩⟩, simp [set.diagonal] } @[simp] lemma range_diag : range (λ x, (x, x)) = diagonal α := by { ext ⟨x, y⟩, simp [diagonal, eq_comm] } +lemma diagonal_subset_iff {s} : diagonal α ⊆ s ↔ ∀ x, (x, x) ∈ s := +by rw [← range_diag, range_subset_iff] + @[simp] lemma prod_subset_compl_diagonal_iff_disjoint : s ×ˢ t ⊆ (diagonal α)ᶜ ↔ disjoint s t := -subset_compl_comm.trans $ by simp_rw [← range_diag, range_subset_iff, - disjoint_left, mem_compl_iff, prod_mk_mem_set_prod_eq, not_and] +prod_subset_iff.trans disjoint_iff_forall_ne.symm @[simp] lemma diag_preimage_prod (s t : set α) : (λ x, (x, x)) ⁻¹' (s ×ˢ t) = s ∩ t := rfl diff --git a/src/order/filter/bases.lean b/src/order/filter/bases.lean index 2b8d36b1ce91d..25eba813bef2b 100644 --- a/src/order/filter/bases.lean +++ b/src/order/filter/bases.lean @@ -347,12 +347,12 @@ begin exact forall_congr (λ s, ⟨λ h, h.1, λ h, ⟨h, λ ⟨t, hl, hP, hts⟩, mem_of_superset hl hts⟩⟩) end -lemma has_basis.comp_of_surjective (h : l.has_basis p s) {g : ι' → ι} (hg : function.surjective g) : +lemma has_basis.comp_surjective (h : l.has_basis p s) {g : ι' → ι} (hg : function.surjective g) : l.has_basis (p ∘ g) (s ∘ g) := ⟨λ t, h.mem_iff.trans hg.exists⟩ lemma has_basis.comp_equiv (h : l.has_basis p s) (e : ι' ≃ ι) : l.has_basis (p ∘ e) (s ∘ e) := -h.comp_of_surjective e.surjective +h.comp_surjective e.surjective /-- If `{s i | p i}` is a basis of a filter `l` and each `s i` includes `s j` such that `p j ∧ q j`, then `{s j | p j ∧ q j}` is a basis of `l`. -/ diff --git a/src/order/filter/basic.lean b/src/order/filter/basic.lean index e98ebb54d07cf..9b280da249ad0 100644 --- a/src/order/filter/basic.lean +++ b/src/order/filter/basic.lean @@ -887,6 +887,8 @@ empty_mem_iff_bot.symm.trans $ mem_principal.trans subset_empty_iff @[simp] lemma principal_ne_bot_iff {s : set α} : ne_bot (𝓟 s) ↔ s.nonempty := ne_bot_iff.trans $ (not_congr principal_eq_bot_iff).trans nonempty_iff_ne_empty.symm +alias principal_ne_bot_iff ↔ _ _root_.set.nonempty.principal_ne_bot + lemma is_compl_principal (s : set α) : is_compl (𝓟 s) (𝓟 sᶜ) := is_compl.of_eq (by rw [inf_principal, inter_compl_self, principal_empty]) $ by rw [sup_principal, union_compl_self, principal_univ] @@ -1622,6 +1624,11 @@ lemma mem_comap' : s ∈ comap f l ↔ {y | ∀ ⦃x⦄, f x = y → x ∈ s} ⟨λ ⟨t, ht, hts⟩, mem_of_superset ht $ λ y hy x hx, hts $ mem_preimage.2 $ by rwa hx, λ h, ⟨_, h, λ x hx, hx rfl⟩⟩ +/-- RHS form is used, e.g., in the definition of `uniform_space`. -/ +lemma mem_comap_prod_mk {x : α} {s : set β} {F : filter (α × β)} : + s ∈ comap (prod.mk x) F ↔ {p : α × β | p.fst = x → p.snd ∈ s} ∈ F := +by simp_rw [mem_comap', prod.ext_iff, and_imp, @forall_swap β (_ = _), forall_eq, eq_comm] + @[simp] lemma eventually_comap : (∀ᶠ a in comap f l, p a) ↔ ∀ᶠ b in l, ∀ a, f a = b → p a := mem_comap' @@ -1728,6 +1735,8 @@ preimage_mem_comap hf lemma comap_id : comap id f = f := le_antisymm (λ s, preimage_mem_comap) (λ s ⟨t, ht, hst⟩, mem_of_superset ht hst) +lemma comap_id' : comap (λ x, x) f = f := comap_id + lemma comap_const_of_not_mem {x : β} (ht : t ∈ g) (hx : x ∉ t) : comap (λ y : α, x) g = ⊥ := empty_mem_iff_bot.1 $ mem_comap'.2 $ mem_of_superset ht $ λ x' hx' y h, hx $ h.symm ▸ hx' diff --git a/src/topology/algebra/uniform_group.lean b/src/topology/algebra/uniform_group.lean index 802b4c71376a9..e731c5798fa64 100644 --- a/src/topology/algebra/uniform_group.lean +++ b/src/topology/algebra/uniform_group.lean @@ -436,7 +436,7 @@ def topological_group.to_uniform_space : uniform_space G := simpa using V_sum _ Hz2 _ Hz1, end, exact set.subset.trans comp_rel_sub U_sub }, - { exact monotone_comp_rel monotone_id monotone_id } + { exact monotone_id.comp_rel monotone_id } end, is_open_uniformity := begin diff --git a/src/topology/uniform_space/basic.lean b/src/topology/uniform_space/basic.lean index c0f70498831b3..6e08e4ec9d8e3 100644 --- a/src/topology/uniform_space/basic.lean +++ b/src/topology/uniform_space/basic.lean @@ -142,7 +142,7 @@ localized "infix (name := uniformity.comp_rel) ` ○ `:55 := comp_rel" in unifor @[simp] theorem swap_id_rel : prod.swap '' id_rel = @id_rel α := set.ext $ assume ⟨a, b⟩, by simp [image_swap_eq_preimage_swap]; exact eq_comm -theorem monotone_comp_rel [preorder β] {f g : β → set (α×α)} +theorem monotone.comp_rel [preorder β] {f g : β → set (α×α)} (hf : monotone f) (hg : monotone g) : monotone (λx, (f x) ○ (g x)) := assume a b h p ⟨z, h₁, h₂⟩, ⟨z, hf h h₁, hg h h₂⟩ @@ -218,12 +218,7 @@ def uniform_space.core.mk' {α : Type u} (U : filter (α × α)) (symm : ∀ r ∈ U, prod.swap ⁻¹' r ∈ U) (comp : ∀ r ∈ U, ∃ t ∈ U, t ○ t ⊆ r) : uniform_space.core α := ⟨U, λ r ru, id_rel_subset.2 (refl _ ru), symm, - begin - intros r ru, - rw [mem_lift'_sets], - exact comp _ ru, - apply monotone_comp_rel; exact monotone_id, - end⟩ + λ r ru, let ⟨s, hs, hsr⟩ := comp _ ru in mem_of_superset (mem_lift' hs) hsr⟩ /-- Defining an `uniform_space.core` from a filter basis satisfying some uniformity-like axioms. -/ def uniform_space.core.mk_of_basis {α : Type u} (B : filter_basis (α × α)) @@ -233,7 +228,7 @@ def uniform_space.core.mk_of_basis {α : Type u} (B : filter_basis (α × α)) { uniformity := B.filter, refl := B.has_basis.ge_iff.mpr (λ r ru, id_rel_subset.2 $ refl _ ru), symm := (B.has_basis.tendsto_iff B.has_basis).mpr symm, - comp := (has_basis.le_basis_iff (B.has_basis.lift' (monotone_comp_rel monotone_id monotone_id)) + comp := (has_basis.le_basis_iff (B.has_basis.lift' (monotone_id.comp_rel monotone_id)) B.has_basis).mpr comp } /-- A uniform space generates a topological space -/ @@ -248,7 +243,7 @@ def uniform_space.core.to_topological_space {α : Type u} (u : uniform_space.cor lemma uniform_space.core_eq : ∀{u₁ u₂ : uniform_space.core α}, u₁.uniformity = u₂.uniformity → u₁ = u₂ -| ⟨u₁, _, _, _⟩ ⟨u₂, _, _, _⟩ h := by { congr, exact h } +| ⟨u₁, _, _, _⟩ ⟨u₂, _, _, _⟩ rfl := by congr -- the topological structure is embedded in the uniform structure -- to avoid instance diamond issues. See Note [forgetful inheritance]. @@ -329,11 +324,7 @@ lemma refl_le_uniformity : 𝓟 id_rel ≤ 𝓤 α := (@uniform_space.to_core α _).refl instance uniformity.ne_bot [nonempty α] : ne_bot (𝓤 α) := -begin - inhabit α, - refine (principal_ne_bot_iff.2 _).mono refl_le_uniformity, - exact ⟨(default, default), rfl⟩ -end +diagonal_nonempty.principal_ne_bot.mono refl_le_uniformity lemma refl_mem_uniformity {x : α} {s : set (α × α)} (h : s ∈ 𝓤 α) : (x, x) ∈ s := @@ -341,7 +332,7 @@ refl_le_uniformity h rfl lemma mem_uniformity_of_eq {x y : α} {s : set (α × α)} (h : s ∈ 𝓤 α) (hx : x = y) : (x, y) ∈ s := -hx ▸ refl_mem_uniformity h +refl_le_uniformity h hx lemma symm_le_uniformity : map (@prod.swap α α) (𝓤 _) ≤ (𝓤 _) := (@uniform_space.to_core α _).symm @@ -356,7 +347,7 @@ lemma comp_mem_uniformity_sets {s : set (α × α)} (hs : s ∈ 𝓤 α) : ∃ t ∈ 𝓤 α, t ○ t ⊆ s := have s ∈ (𝓤 α).lift' (λt:set (α×α), t ○ t), from comp_le_uniformity hs, -(mem_lift'_sets $ monotone_comp_rel monotone_id monotone_id).mp this +(mem_lift'_sets $ monotone_id.comp_rel monotone_id).mp this /-- If `s ∈ 𝓤 α`, then for any natural `n`, for a subset `t` of a sufficiently small set in `𝓤 α`, we have `t ○ t ○ ... ○ t ⊆ s` (`n` compositions). -/ @@ -411,7 +402,7 @@ lemma comp_symm_of_uniformity {s : set (α × α)} (hs : s ∈ 𝓤 α) : ∃ t ∈ 𝓤 α, (∀{a b}, (a, b) ∈ t → (b, a) ∈ t) ∧ t ○ t ⊆ s := let ⟨t, ht₁, ht₂⟩ := comp_mem_uniformity_sets hs in let ⟨t', ht', ht'₁, ht'₂⟩ := symm_of_uniformity ht₁ in -⟨t', ht', ht'₁, subset.trans (monotone_comp_rel monotone_id monotone_id ht'₂) ht₂⟩ +⟨t', ht', ht'₁, subset.trans (monotone_id.comp_rel monotone_id ht'₂) ht₂⟩ lemma uniformity_le_symm : 𝓤 α ≤ (@prod.swap α α) <$> 𝓤 α := by rw [map_swap_eq_comap_swap]; @@ -449,7 +440,7 @@ calc (𝓤 α).lift (λs, f (s ○ s)) = ((𝓤 α).lift' (λs:set (α×α), s ○ s)).lift f : begin rw [lift_lift'_assoc], - exact monotone_comp_rel monotone_id monotone_id, + exact monotone_id.comp_rel monotone_id, exact h end ... ≤ (𝓤 α).lift f : lift_mono comp_le_uniformity le_rfl @@ -460,16 +451,16 @@ calc (𝓤 α).lift' (λd, d ○ (d ○ d)) = (𝓤 α).lift (λs, (𝓤 α).lift' (λt:set(α×α), s ○ (t ○ t))) : begin rw [lift_lift'_same_eq_lift'], - exact (assume x, monotone_comp_rel monotone_const $ monotone_comp_rel monotone_id monotone_id), - exact (assume x, monotone_comp_rel monotone_id monotone_const), + exact (assume x, monotone_const.comp_rel $ monotone_id.comp_rel monotone_id), + exact (assume x, monotone_id.comp_rel monotone_const), end ... ≤ (𝓤 α).lift (λs, (𝓤 α).lift' (λt:set(α×α), s ○ t)) : lift_mono' $ assume s hs, @uniformity_lift_le_comp α _ _ (𝓟 ∘ (○) s) $ - monotone_principal.comp (monotone_comp_rel monotone_const monotone_id) + monotone_principal.comp (monotone_const.comp_rel monotone_id) ... = (𝓤 α).lift' (λs:set(α×α), s ○ s) : lift_lift'_same_eq_lift' - (assume s, monotone_comp_rel monotone_const monotone_id) - (assume s, monotone_comp_rel monotone_id monotone_const) + (assume s, monotone_const.comp_rel monotone_id) + (assume s, monotone_id.comp_rel monotone_const) ... ≤ (𝓤 α) : comp_le_uniformity /-- See also `comp_open_symm_mem_uniformity_sets`. -/ @@ -591,15 +582,8 @@ lemma mem_nhds_uniformity_iff_left {x : α} {s : set α} : s ∈ 𝓝 x ↔ {p : α × α | p.2 = x → p.1 ∈ s} ∈ 𝓤 α := by { rw [uniformity_eq_symm, mem_nhds_uniformity_iff_right], refl } -lemma nhds_eq_comap_uniformity_aux {α : Type u} {x : α} {s : set α} {F : filter (α × α)} : - {p : α × α | p.fst = x → p.snd ∈ s} ∈ F ↔ s ∈ comap (prod.mk x) F := -by rw mem_comap ; from iff.intro - (assume hs, ⟨_, hs, assume x hx, hx rfl⟩) - (assume ⟨t, h, ht⟩, F.sets_of_superset h $ - assume ⟨p₁, p₂⟩ hp (h : p₁ = x), ht $ by simp [h.symm, hp]) - lemma nhds_eq_comap_uniformity {x : α} : 𝓝 x = (𝓤 α).comap (prod.mk x) := -by { ext s, rw [mem_nhds_uniformity_iff_right], exact nhds_eq_comap_uniformity_aux } +by { ext s, rw [mem_nhds_uniformity_iff_right, mem_comap_prod_mk] } /-- See also `is_open_iff_open_ball_subset`. -/ lemma is_open_iff_ball_subset {s : set α} : is_open s ↔ ∀ x ∈ s, ∃ V ∈ 𝓤 α, ball x V ⊆ s := @@ -621,6 +605,9 @@ begin exact nhds_basis_uniformity' h end +lemma nhds_eq_comap_uniformity' {x : α} : 𝓝 x = (𝓤 α).comap (λ y, (y, x)) := +(nhds_basis_uniformity (𝓤 α).basis_sets).eq_of_same_basis $ (𝓤 α).basis_sets.comap _ + lemma uniform_space.mem_nhds_iff {x : α} {s : set α} : s ∈ 𝓝 x ↔ ∃ V ∈ 𝓤 α, ball x V ⊆ s := begin rw [nhds_eq_comap_uniformity, mem_comap], @@ -747,22 +734,12 @@ lemma tendsto_left_nhds_uniformity {a : α} : tendsto (λa', (a, a')) (𝓝 a) ( assume s, mem_nhds_left a lemma lift_nhds_left {x : α} {g : set α → filter β} (hg : monotone g) : - (𝓝 x).lift g = (𝓤 α).lift (λs:set (α×α), g {y | (x, y) ∈ s}) := -eq.trans - begin - rw [nhds_eq_uniformity], - exact (filter.lift_assoc $ monotone_principal.comp $ monotone_preimage.comp monotone_preimage ) - end - (congr_arg _ $ funext $ assume s, filter.lift_principal hg) + (𝓝 x).lift g = (𝓤 α).lift (λs:set (α×α), g (ball x s)) := +by { rw [nhds_eq_comap_uniformity, comap_lift_eq2 hg], refl } lemma lift_nhds_right {x : α} {g : set α → filter β} (hg : monotone g) : (𝓝 x).lift g = (𝓤 α).lift (λs:set (α×α), g {y | (y, x) ∈ s}) := -calc (𝓝 x).lift g = (𝓤 α).lift (λs:set (α×α), g {y | (x, y) ∈ s}) : lift_nhds_left hg - ... = ((@prod.swap α α) <$> (𝓤 α)).lift (λs:set (α×α), g {y | (x, y) ∈ s}) : - by rw [←uniformity_eq_symm] - ... = (𝓤 α).lift (λs:set (α×α), g {y | (x, y) ∈ image prod.swap s}) : - map_lift_eq2 $ hg.comp monotone_preimage - ... = _ : by simp [image_swap_eq_preimage_swap] +by { rw [nhds_eq_comap_uniformity', comap_lift_eq2 hg], refl } lemma nhds_nhds_eq_uniformity_uniformity_prod {a b : α} : 𝓝 a ×ᶠ 𝓝 b = @@ -770,9 +747,7 @@ lemma nhds_nhds_eq_uniformity_uniformity_prod {a b : α} : {y : α | (y, a) ∈ s} ×ˢ {y : α | (b, y) ∈ t})) := begin rw [nhds_eq_uniformity', nhds_eq_uniformity, prod_lift'_lift'], - { refl }, - { exact monotone_preimage }, - { exact monotone_preimage }, + exacts [rfl, monotone_preimage, monotone_preimage] end lemma nhds_eq_uniformity_prod {a b : α} : @@ -874,7 +849,7 @@ lemma uniformity_eq_uniformity_interior : 𝓤 α = (𝓤 α).lift' interior := le_antisymm (le_infi $ assume d, le_infi $ assume hd, let ⟨s, hs, hs_comp⟩ := (mem_lift'_sets $ - monotone_comp_rel monotone_id $ monotone_comp_rel monotone_id monotone_id).mp + monotone_id.comp_rel $ monotone_id.comp_rel monotone_id).mp (comp_le_uniformity3 hd) in let ⟨t, ht, hst, ht_comp⟩ := nhdset_of_mem_uniformity s hs in have s ⊆ interior d, from @@ -1062,12 +1037,8 @@ instance : has_bot (uniform_space α) := ⟨{ to_topological_space := ⊥, uniformity := 𝓟 id_rel, refl := le_rfl, - symm := by simp [tendsto]; apply subset.refl, - comp := - begin - rw [lift'_principal], {simp}, - exact monotone_comp_rel monotone_id monotone_id - end, + symm := by simp [tendsto], + comp := lift'_le (mem_principal_self _) $ principal_mono.2 id_comp_rel.subset, is_open_uniformity := assume s, by simp [is_open_fold, subset_def, id_rel] {contextual := tt } } ⟩ @@ -1138,19 +1109,11 @@ def uniform_space.comap (f : α → β) (u : uniform_space β) : uniform_space begin rw [comap_lift'_eq, comap_lift'_eq2], exact (lift'_mono' $ assume s hs ⟨a₁, a₂⟩ ⟨x, h₁, h₂⟩, ⟨f x, h₁, h₂⟩), - exact monotone_comp_rel monotone_id monotone_id + exact monotone_id.comp_rel monotone_id end (comap_mono u.comp), - is_open_uniformity := λ s, begin - change (@is_open α (u.to_topological_space.induced f) s ↔ _), - simp [is_open_iff_nhds, nhds_induced, mem_nhds_uniformity_iff_right, filter.comap, and_comm], - refine ball_congr (λ x hx, ⟨_, _⟩), - { rintro ⟨t, hts, ht⟩, refine ⟨_, ht, _⟩, - rintro ⟨x₁, x₂⟩ h rfl, exact hts (h rfl) }, - { rintro ⟨t, ht, hts⟩, - exact ⟨{y | (f x, y) ∈ t}, λ y hy, @hts (x, y) hy rfl, - mem_nhds_uniformity_iff_right.1 $ mem_nhds_left _ ht⟩ } - end } + is_open_uniformity := λ s, by simp only [is_open_fold, is_open_induced, is_open_iff_mem_nhds, + nhds_induced, nhds_eq_comap_uniformity, comap_comap, ← mem_comap_prod_mk, ← uniformity] } lemma uniformity_comap [uniform_space α] [uniform_space β] {f : α → β} (h : ‹uniform_space α› = uniform_space.comap f ‹uniform_space β›) : @@ -1655,9 +1618,6 @@ end sum end constructions --- For a version of the Lebesgue number lemma assuming only a sequentially compact space, --- see topology/sequences.lean - /-- Let `c : ι → set α` be an open cover of a compact set `s`. Then there exists an entourage `n` such that for each `x ∈ s` its `n`-neighborhood is contained in some `c i`. -/ lemma lebesgue_number_lemma {α : Type u} [uniform_space α] {s : set α} {ι} {c : ι → set α} @@ -1671,7 +1631,7 @@ begin rcases comp_mem_uniformity_sets hm with ⟨m', hm', mm'⟩, apply (𝓤 α).sets_of_superset hm', rintros ⟨x, y⟩ hp rfl, - refine ⟨i, m', hm', λ z hz, h (monotone_comp_rel monotone_id monotone_const mm' _)⟩, + refine ⟨i, m', hm', λ z hz, h (monotone_id.comp_rel monotone_const mm' _)⟩, dsimp [-mem_comp_rel] at hz ⊢, rw comp_rel_assoc, exact ⟨y, hp, hz⟩ }, have hu₂ : s ⊆ ⋃ n ∈ 𝓤 α, u n, @@ -1741,13 +1701,11 @@ variables [uniform_space α] theorem tendsto_nhds_right {f : filter β} {u : β → α} {a : α} : tendsto u f (𝓝 a) ↔ tendsto (λ x, (a, u x)) f (𝓤 α) := -⟨λ H, tendsto_left_nhds_uniformity.comp H, -λ H s hs, by simpa [mem_of_mem_nhds hs] using H (mem_nhds_uniformity_iff_right.1 hs)⟩ +by rw [nhds_eq_comap_uniformity, tendsto_comap_iff] theorem tendsto_nhds_left {f : filter β} {u : β → α} {a : α} : tendsto u f (𝓝 a) ↔ tendsto (λ x, (u x, a)) f (𝓤 α) := -⟨λ H, tendsto_right_nhds_uniformity.comp H, -λ H s hs, by simpa [mem_of_mem_nhds hs] using H (mem_nhds_uniformity_iff_left.1 hs)⟩ +by rw [nhds_eq_comap_uniformity', tendsto_comap_iff] theorem continuous_at_iff'_right [topological_space β] {f : β → α} {b : β} : continuous_at f b ↔ tendsto (λ x, (f b, f x)) (𝓝 b) (𝓤 α) := diff --git a/src/topology/uniform_space/compact.lean b/src/topology/uniform_space/compact.lean index 68806ee7d7932..75744f7b46199 100644 --- a/src/topology/uniform_space/compact.lean +++ b/src/topology/uniform_space/compact.lean @@ -73,24 +73,15 @@ end /-- The unique uniform structure inducing a given compact topological structure. -/ def uniform_space_of_compact_t2 [topological_space γ] [compact_space γ] [t2_space γ] : uniform_space γ := -{ uniformity := ⨆ x, 𝓝 (x, x), - refl := begin - simp_rw [filter.principal_le_iff, mem_supr], - rintros V V_in ⟨x, _⟩ ⟨⟩, - exact mem_of_mem_nhds (V_in x), - end, - symm := begin - refine le_of_eq _, - rw filter.map_supr, - congr' with x : 1, - erw [nhds_prod_eq, ← prod_comm], - end, +{ uniformity := 𝓝ˢ (diagonal γ), + refl := principal_le_nhds_set, + symm := continuous_swap.tendsto_nhds_set $ λ x, eq.symm, comp := begin /- - This is the difficult part of the proof. We need to prove that, for each neighborhood W - of the diagonal Δ, W ○ W is still a neighborhood of the diagonal. + This is the difficult part of the proof. We need to prove that, for each neighborhood `W` + of the diagonal `Δ`, there exists a smaller neighborhood `V` such that `V ○ V ⊆ W`. -/ - set 𝓝Δ := ⨆ x : γ, 𝓝 (x, x), -- The filter of neighborhoods of Δ + set 𝓝Δ := 𝓝ˢ (diagonal γ), -- The filter of neighborhoods of Δ set F := 𝓝Δ.lift' (λ (s : set (γ × γ)), s ○ s), -- Compositions of neighborhoods of Δ -- If this weren't true, then there would be V ∈ 𝓝Δ such that F ⊓ 𝓟 Vᶜ ≠ ⊥ rw le_iff_forall_inf_principal_compl, @@ -106,14 +97,9 @@ def uniform_space_of_compact_t2 [topological_space γ] [compact_space γ] [t2_sp { have : (x, y) ∈ closure (Vᶜ), by rwa mem_closure_iff_cluster_pt, rwa closure_compl at this }, have diag_subset : diagonal γ ⊆ interior V, - { rw subset_interior_iff_nhds, - rintros ⟨x, x⟩ ⟨⟩, - exact (mem_supr.mp V_in : _) x }, + from subset_interior_iff_mem_nhds_set.2 V_in, have x_ne_y : x ≠ y, - { intro h, - apply this, - apply diag_subset, - simp [h] }, + from mt (@diag_subset (x, y)) this, -- Since γ is compact and Hausdorff, it is normal, hence T₃. haveI : normal_space γ := normal_of_compact_t2, -- So there are closed neighboords V₁ and V₂ of x and y contained in disjoint open neighborhoods @@ -124,20 +110,15 @@ def uniform_space_of_compact_t2 [topological_space γ] [compact_space γ] [t2_sp -- We set U₃ := (V₁ ∪ V₂)ᶜ so that W := U₁ ×ˢ U₁ ∪ U₂ ×ˢ U₂ ∪ U₃ ×ˢ U₃ is an open -- neighborhood of Δ. let U₃ := (V₁ ∪ V₂)ᶜ, - have U₃_op : is_open U₃ := - is_open_compl_iff.mpr (is_closed.union V₁_cl V₂_cl), + have U₃_op : is_open U₃ := (V₁_cl.union V₂_cl).is_open_compl, let W := U₁ ×ˢ U₁ ∪ U₂ ×ˢ U₂ ∪ U₃ ×ˢ U₃, have W_in : W ∈ 𝓝Δ, - { rw mem_supr, - intros x, - apply is_open.mem_nhds (is_open.union (is_open.union _ _) _), - { by_cases hx : x ∈ V₁ ∪ V₂, - { left, - cases hx with hx hx ; [left, right] ; split ; tauto }, - { right, - rw mem_prod, - tauto }, }, - all_goals { simp only [is_open.prod, *] } }, + { rw [mem_nhds_set_iff_forall], + rintros ⟨z, z'⟩ (rfl : z = z'), + refine is_open.mem_nhds _ _, + { apply_rules [is_open.union, is_open.prod] }, + { simp only [mem_union, mem_prod, and_self], + exact (em _).imp_left (λ h, union_subset_union VU₁ VU₂ h) } }, -- So W ○ W ∈ F by definition of F have : W ○ W ∈ F, by simpa only using mem_lift' W_in, -- And V₁ ×ˢ V₂ ∈ 𝓝 (x, y) @@ -160,19 +141,16 @@ def uniform_space_of_compact_t2 [topological_space γ] [compact_space γ] [t2_sp is_open_uniformity := begin -- Here we need to prove the topology induced by the constructed uniformity is the -- topology we started with. - suffices : ∀ x : γ, filter.comap (prod.mk x) (⨆ y, 𝓝 (y ,y)) = 𝓝 x, + suffices : ∀ x : γ, filter.comap (prod.mk x) (𝓝ˢ (diagonal γ)) = 𝓝 x, { intros s, - change is_open s ↔ _, - simp_rw [is_open_iff_mem_nhds, nhds_eq_comap_uniformity_aux, this] }, + simp_rw [is_open_fold, is_open_iff_mem_nhds, ← mem_comap_prod_mk, this] }, intros x, - simp_rw [comap_supr, nhds_prod_eq, comap_prod, - show prod.fst ∘ prod.mk x = λ y : γ, x, by ext ; simp, - show prod.snd ∘ (prod.mk x) = (id : γ → γ), by ext ; refl, comap_id], + simp_rw [nhds_set_diagonal, comap_supr, nhds_prod_eq, comap_prod, (∘), comap_id'], rw [supr_split_single _ x, comap_const_of_mem (λ V, mem_of_mem_nhds)], suffices : ∀ y ≠ x, comap (λ (y : γ), x) (𝓝 y) ⊓ 𝓝 y ≤ 𝓝 x, by simpa, intros y hxy, - simp [comap_const_of_not_mem (compl_singleton_mem_nhds hxy) (by simp)], + simp [comap_const_of_not_mem (compl_singleton_mem_nhds hxy) (not_not.2 rfl)] end } /-! @@ -183,12 +161,9 @@ def uniform_space_of_compact_t2 [topological_space γ] [compact_space γ] [t2_sp continuous. -/ lemma compact_space.uniform_continuous_of_continuous [compact_space α] {f : α → β} (h : continuous f) : uniform_continuous f := -calc -map (prod.map f f) (𝓤 α) = map (prod.map f f) (⨆ x, 𝓝 (x, x)) : by rw compact_space_uniformity - ... = ⨆ x, map (prod.map f f) (𝓝 (x, x)) : by rw filter.map_supr - ... ≤ ⨆ x, 𝓝 (f x, f x) : supr_mono (λ x, (h.prod_map h).continuous_at) - ... ≤ ⨆ y, 𝓝 (y, y) : supr_comp_le (λ y, 𝓝 (y, y)) f - ... ≤ 𝓤 β : supr_nhds_le_uniformity +have tendsto (prod.map f f) (𝓝ˢ (diagonal α)) (𝓝ˢ (diagonal β)), + from (h.prod_map h).tendsto_nhds_set maps_to_prod_map_diagonal, +(this.mono_left nhds_set_diagonal_eq_uniformity.ge).mono_right nhds_set_diagonal_le_uniformity /-- Heine-Cantor: a continuous function on a compact set of a uniform space is uniformly continuous. -/ diff --git a/src/topology/uniform_space/completion.lean b/src/topology/uniform_space/completion.lean index ae7133be2cc56..2158ae9f3b529 100644 --- a/src/topology/uniform_space/completion.lean +++ b/src/topology/uniform_space/completion.lean @@ -111,14 +111,14 @@ calc ((𝓤 α).lift' gen).lift' (λs, comp_rel s s) = begin rw [lift'_lift'_assoc], exact monotone_gen, - exact (monotone_comp_rel monotone_id monotone_id) + exact monotone_id.comp_rel monotone_id end ... ≤ (𝓤 α).lift' (λs, gen $ comp_rel s s) : lift'_mono' $ assume s hs, comp_rel_gen_gen_subset_gen_comp_rel ... = ((𝓤 α).lift' $ λs:set(α×α), comp_rel s s).lift' gen : begin rw [lift'_lift'_assoc], - exact (monotone_comp_rel monotone_id monotone_id), + exact monotone_id.comp_rel monotone_id, exact monotone_gen end ... ≤ (𝓤 α).lift' gen : lift'_mono comp_le_uniformity le_rfl diff --git a/src/topology/uniform_space/separation.lean b/src/topology/uniform_space/separation.lean index 778166fc29fff..39c1ea89295cc 100644 --- a/src/topology/uniform_space/separation.lean +++ b/src/topology/uniform_space/separation.lean @@ -244,7 +244,7 @@ instance separation_setoid.uniform_space {α : Type u} [u : uniform_space α] : by simp [prod.swap, (∘)]; exact tendsto_map.comp tendsto_swap_uniformity, comp := calc (map (λ (p : α × α), (⟦p.fst⟧, ⟦p.snd⟧)) u.uniformity).lift' (λs, comp_rel s s) = u.uniformity.lift' ((λs, comp_rel s s) ∘ image (λ (p : α × α), (⟦p.fst⟧, ⟦p.snd⟧))) : - map_lift'_eq2 $ monotone_comp_rel monotone_id monotone_id + map_lift'_eq2 $ monotone_id.comp_rel monotone_id ... ≤ u.uniformity.lift' (image (λ (p : α × α), (⟦p.fst⟧, ⟦p.snd⟧)) ∘ (λs:set (α×α), comp_rel s (comp_rel s s))) : lift'_mono' $ assume s hs ⟨a, b⟩ ⟨c, ⟨⟨a₁, a₂⟩, ha, a_eq⟩, ⟨⟨b₁, b₂⟩, hb, b_eq⟩⟩, @@ -259,7 +259,7 @@ instance separation_setoid.uniform_space {α : Type u} [u : uniform_space α] : ... = map (λp:(α×α), (⟦p.1⟧, ⟦p.2⟧)) (u.uniformity.lift' (λs:set (α×α), comp_rel s (comp_rel s s))) : by rw [map_lift'_eq]; - exact monotone_comp_rel monotone_id (monotone_comp_rel monotone_id monotone_id) + exact monotone_id.comp_rel (monotone_id.comp_rel monotone_id) ... ≤ map (λp:(α×α), (⟦p.1⟧, ⟦p.2⟧)) u.uniformity : map_mono comp_le_uniformity3, is_open_uniformity := assume s, diff --git a/src/topology/uniform_space/uniform_embedding.lean b/src/topology/uniform_space/uniform_embedding.lean index c2871bd762278..c38a4deeafc8a 100644 --- a/src/topology/uniform_space/uniform_embedding.lean +++ b/src/topology/uniform_space/uniform_embedding.lean @@ -511,7 +511,7 @@ by simpa only [dense_inducing.extend] using tendsto_nhds_lim (uniformly_extend_e lemma uniform_continuous_uniformly_extend [cγ : complete_space γ] : uniform_continuous ψ := assume d hd, let ⟨s, hs, hs_comp⟩ := (mem_lift'_sets $ - monotone_comp_rel monotone_id $ monotone_comp_rel monotone_id monotone_id).mp + monotone_id.comp_rel $ monotone_id.comp_rel monotone_id).mp (comp_le_uniformity3 hd) in have h_pnt : ∀{a m}, m ∈ 𝓝 a → ∃c, c ∈ f '' preimage e m ∧ (c, ψ a) ∈ s ∧ (ψ a, c) ∈ s, from assume a m hm, From e7e2ba8aa216a5833b5ed85a93317263711a36b5 Mon Sep 17 00:00:00 2001 From: negiizhao Date: Wed, 25 Jan 2023 11:06:52 +0000 Subject: [PATCH 0348/1291] feat(algebra/order/ring/with_top): ring covariance & ordered ring typeclasses for `with_bot` (#18149) mathlib4 PR: https://github.com/leanprover-community/mathlib4/pull/1508 --- src/algebra/order/monoid/with_top.lean | 18 +++- src/algebra/order/ring/with_top.lean | 123 ++++++++++++++++++++++--- src/data/nat/with_bot.lean | 4 +- 3 files changed, 127 insertions(+), 18 deletions(-) diff --git a/src/algebra/order/monoid/with_top.lean b/src/algebra/order/monoid/with_top.lean index 2d35a1f9616ef..b80bfa8232f70 100644 --- a/src/algebra/order/monoid/with_top.lean +++ b/src/algebra/order/monoid/with_top.lean @@ -31,6 +31,12 @@ variables [has_one α] @[simp, norm_cast, to_additive] lemma coe_eq_one {a : α} : (a : with_top α) = 1 ↔ a = 1 := coe_eq_coe +@[simp, norm_cast, to_additive coe_nonneg] +lemma one_le_coe [has_le α] {a : α} : 1 ≤ (a : with_top α) ↔ 1 ≤ a := coe_le_coe + +@[simp, norm_cast, to_additive coe_le_zero] +lemma coe_le_one [has_le α] {a : α} : (a : with_top α) ≤ 1 ↔ a ≤ 1 := coe_le_coe + @[simp, norm_cast, to_additive coe_pos] lemma one_lt_coe [has_lt α] {a : α} : 1 < (a : with_top α) ↔ 1 < a := coe_lt_coe @@ -341,13 +347,19 @@ lemma coe_one [has_one α] : ((1 : α) : with_bot α) = 1 := rfl lemma coe_eq_one [has_one α] {a : α} : (a : with_bot α) = 1 ↔ a = 1 := with_top.coe_eq_one -@[norm_cast, to_additive coe_pos] +@[simp, norm_cast, to_additive coe_nonneg] +lemma one_le_coe [has_one α] [has_le α] {a : α} : 1 ≤ (a : with_bot α) ↔ 1 ≤ a := coe_le_coe + +@[simp, norm_cast, to_additive coe_le_zero] +lemma coe_le_one [has_one α] [has_le α] {a : α} : (a : with_bot α) ≤ 1 ↔ a ≤ 1 := coe_le_coe + +@[simp, norm_cast, to_additive coe_pos] lemma one_lt_coe [has_one α] [has_lt α] {a : α} : 1 < (a : with_bot α) ↔ 1 < a := coe_lt_coe -@[norm_cast, to_additive coe_lt_zero] +@[simp, norm_cast, to_additive coe_lt_zero] lemma coe_lt_one [has_one α] [has_lt α] {a : α} : (a : with_bot α) < 1 ↔ a < 1 := coe_lt_coe -@[to_additive] protected lemma map_one {β} [has_one α] (f : α → β) : +@[simp, to_additive] protected lemma map_one {β} [has_one α] (f : α → β) : (1 : with_bot α).map f = (f 1 : with_bot β) := rfl @[norm_cast] lemma coe_nat [add_monoid_with_one α] (n : ℕ) : ((n : α) : with_bot α) = n := rfl diff --git a/src/algebra/order/ring/with_top.lean b/src/algebra/order/ring/with_top.lean index b67c04b9a94ce..3cc33489d4e66 100644 --- a/src/algebra/order/ring/with_top.lean +++ b/src/algebra/order/ring/with_top.lean @@ -13,7 +13,7 @@ import algebra.order.ring.canonical > Any changes to this file require a corresponding PR to mathlib4. The main results of this section are `with_top.canonically_ordered_comm_semiring` and -`with_bot.comm_monoid_with_zero`. +`with_bot.ordered_comm_semiring`. -/ variables {α : Type*} @@ -260,17 +260,114 @@ with_top.comm_monoid_with_zero instance [canonically_ordered_comm_semiring α] [nontrivial α] : comm_semiring (with_bot α) := with_top.comm_semiring -instance [canonically_ordered_comm_semiring α] [nontrivial α] : pos_mul_mono (with_bot α) := -pos_mul_mono_iff_covariant_pos.2 ⟨begin - rintros ⟨x, x0⟩ a b h, simp only [subtype.coe_mk], - lift x to α using x0.ne_bot, - induction a using with_bot.rec_bot_coe, { simp_rw [mul_bot x0.ne.symm, bot_le] }, - induction b using with_bot.rec_bot_coe, { exact absurd h (bot_lt_coe a).not_le }, - simp only [← coe_mul, coe_le_coe] at *, - exact mul_le_mul_left' h x, - end ⟩ - -instance [canonically_ordered_comm_semiring α] [nontrivial α] : mul_pos_mono (with_bot α) := -pos_mul_mono_iff_mul_pos_mono.mp infer_instance +instance [mul_zero_class α] [preorder α] [pos_mul_mono α] : + pos_mul_mono (with_bot α) := +⟨begin + rintros ⟨x, x0⟩ a b h, simp only [subtype.coe_mk], + rcases eq_or_ne x 0 with rfl | x0', { simp, }, + lift x to α, { rintro ⟨rfl⟩, exact (with_bot.bot_lt_coe (0 : α)).not_le x0, }, + induction a using with_bot.rec_bot_coe, { simp_rw [mul_bot x0', bot_le] }, + induction b using with_bot.rec_bot_coe, { exact absurd h (bot_lt_coe a).not_le }, + simp only [← coe_mul, coe_le_coe] at *, + norm_cast at x0, + exact mul_le_mul_of_nonneg_left h x0, +end ⟩ + +instance [mul_zero_class α] [preorder α] [mul_pos_mono α] : + mul_pos_mono (with_bot α) := +⟨begin + rintros ⟨x, x0⟩ a b h, simp only [subtype.coe_mk], + rcases eq_or_ne x 0 with rfl | x0', { simp, }, + lift x to α, { rintro ⟨rfl⟩, exact (with_bot.bot_lt_coe (0 : α)).not_le x0, }, + induction a using with_bot.rec_bot_coe, { simp_rw [bot_mul x0', bot_le] }, + induction b using with_bot.rec_bot_coe, { exact absurd h (bot_lt_coe a).not_le }, + simp only [← coe_mul, coe_le_coe] at *, + norm_cast at x0, + exact mul_le_mul_of_nonneg_right h x0, +end ⟩ + +instance [mul_zero_class α] [preorder α] [pos_mul_strict_mono α] : + pos_mul_strict_mono (with_bot α) := +⟨begin + rintros ⟨x, x0⟩ a b h, simp only [subtype.coe_mk], + lift x to α using x0.ne_bot, + induction b using with_bot.rec_bot_coe, { exact absurd h not_lt_bot, }, + induction a using with_bot.rec_bot_coe, { simp_rw [mul_bot x0.ne.symm, ← coe_mul, bot_lt_coe], }, + simp only [← coe_mul, coe_lt_coe] at *, + norm_cast at x0, + exact mul_lt_mul_of_pos_left h x0, +end ⟩ + +instance [mul_zero_class α] [preorder α] [mul_pos_strict_mono α] : + mul_pos_strict_mono (with_bot α) := +⟨begin + rintros ⟨x, x0⟩ a b h, simp only [subtype.coe_mk], + lift x to α using x0.ne_bot, + induction b using with_bot.rec_bot_coe, { exact absurd h not_lt_bot, }, + induction a using with_bot.rec_bot_coe, { simp_rw [bot_mul x0.ne.symm, ← coe_mul, bot_lt_coe], }, + simp only [← coe_mul, coe_lt_coe] at *, + norm_cast at x0, + exact mul_lt_mul_of_pos_right h x0, +end ⟩ + +instance [mul_zero_class α] [preorder α] [pos_mul_reflect_lt α] : + pos_mul_reflect_lt (with_bot α) := +⟨begin + rintros ⟨x, x0⟩ a b h, simp only [subtype.coe_mk] at h, + rcases eq_or_ne x 0 with rfl | x0', { simpa using h, }, + lift x to α, { rintro ⟨rfl⟩, exact (with_bot.bot_lt_coe (0 : α)).not_le x0, }, + induction b using with_bot.rec_bot_coe, { rw [mul_bot x0'] at h, exact absurd h bot_le.not_lt, }, + induction a using with_bot.rec_bot_coe, { exact with_bot.bot_lt_coe _, }, + simp only [← coe_mul, coe_lt_coe] at *, + norm_cast at x0, + exact lt_of_mul_lt_mul_left h x0, +end ⟩ + +instance [mul_zero_class α] [preorder α] [mul_pos_reflect_lt α] : + mul_pos_reflect_lt (with_bot α) := +⟨begin + rintros ⟨x, x0⟩ a b h, simp only [subtype.coe_mk] at h, + rcases eq_or_ne x 0 with rfl | x0', { simpa using h, }, + lift x to α, { rintro ⟨rfl⟩, exact (with_bot.bot_lt_coe (0 : α)).not_le x0, }, + induction b using with_bot.rec_bot_coe, { rw [bot_mul x0'] at h, exact absurd h bot_le.not_lt, }, + induction a using with_bot.rec_bot_coe, { exact with_bot.bot_lt_coe _, }, + simp only [← coe_mul, coe_lt_coe] at *, + norm_cast at x0, + exact lt_of_mul_lt_mul_right h x0, +end ⟩ + +instance [mul_zero_class α] [preorder α] [pos_mul_mono_rev α] : + pos_mul_mono_rev (with_bot α) := +⟨begin + rintros ⟨x, x0⟩ a b h, simp only [subtype.coe_mk] at h, + lift x to α using x0.ne_bot, + induction a using with_bot.rec_bot_coe, { exact bot_le, }, + induction b using with_bot.rec_bot_coe, + { rw [mul_bot x0.ne.symm, ← coe_mul] at h, exact absurd h (bot_lt_coe (x * a)).not_le, }, + simp only [← coe_mul, coe_le_coe] at *, + norm_cast at x0, + exact le_of_mul_le_mul_left h x0, +end ⟩ + +instance [mul_zero_class α] [preorder α] [mul_pos_mono_rev α] : + mul_pos_mono_rev (with_bot α) := +⟨begin + rintros ⟨x, x0⟩ a b h, simp only [subtype.coe_mk] at h, + lift x to α using x0.ne_bot, + induction a using with_bot.rec_bot_coe, { exact bot_le, }, + induction b using with_bot.rec_bot_coe, + { rw [bot_mul x0.ne.symm, ← coe_mul] at h, exact absurd h (bot_lt_coe (a * x)).not_le, }, + simp only [← coe_mul, coe_le_coe] at *, + norm_cast at x0, + exact le_of_mul_le_mul_right h x0, +end ⟩ + +instance [canonically_ordered_comm_semiring α] [nontrivial α] : + ordered_comm_semiring (with_bot α) := +{ mul_le_mul_of_nonneg_left := λ _ _ _, mul_le_mul_of_nonneg_left, + mul_le_mul_of_nonneg_right := λ _ _ _, mul_le_mul_of_nonneg_right, + .. with_bot.zero_le_one_class, + .. with_bot.ordered_add_comm_monoid, + .. with_bot.comm_semiring, } end with_bot diff --git a/src/data/nat/with_bot.lean b/src/data/nat/with_bot.lean index 713a12e2cc1f8..e98aa714ee23a 100644 --- a/src/data/nat/with_bot.lean +++ b/src/data/nat/with_bot.lean @@ -53,12 +53,12 @@ begin exact add_eq_three_iff end -@[simp] lemma coe_nonneg {n : ℕ} : 0 ≤ (n : with_bot ℕ) := +lemma coe_nonneg {n : ℕ} : 0 ≤ (n : with_bot ℕ) := by { rw [← with_bot.coe_zero, with_bot.coe_le_coe], exact nat.zero_le _ } @[simp] lemma lt_zero_iff (n : with_bot ℕ) : n < 0 ↔ n = ⊥ := option.cases_on n dec_trivial $ λ n, iff_of_false - (by simp [with_bot.some_eq_coe]) (λ h, option.no_confusion h) + (by simp [with_bot.some_eq_coe, coe_nonneg]) (λ h, option.no_confusion h) lemma one_le_iff_zero_lt {x : with_bot ℕ} : 1 ≤ x ↔ 0 < x := begin From 12665d3612d46b1838ec1273d291c21a61fc1707 Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Wed, 25 Jan 2023 11:06:53 +0000 Subject: [PATCH 0349/1291] =?UTF-8?q?feat(data/nat/with=5Fbot):=20`n=20 Date: Wed, 25 Jan 2023 11:06:54 +0000 Subject: [PATCH 0350/1291] feat(data/pequiv,order/initial_seg): use `fun_like`, `embedding_like` (#18198) This is a backport of leanprover-community/mathlib4#1488. --- src/data/pequiv.lean | 30 +++++++++++------------------- src/order/initial_seg.lean | 34 ++++++++++++++++++++++------------ 2 files changed, 33 insertions(+), 31 deletions(-) diff --git a/src/data/pequiv.lean b/src/data/pequiv.lean index 1d2c3b745b9e4..bb10920cf8b6c 100644 --- a/src/data/pequiv.lean +++ b/src/data/pequiv.lean @@ -56,30 +56,22 @@ namespace pequiv variables {α : Type u} {β : Type v} {γ : Type w} {δ : Type x} open function option -instance : has_coe_to_fun (α ≃. β) (λ _, α → option β) := ⟨to_fun⟩ +instance fun_like : fun_like (α ≃. β) α (λ _, option β) := +{ coe := to_fun, + coe_injective' := + begin + rintro ⟨f₁, f₂, hf⟩ ⟨g₁, g₂, hg⟩ (rfl : f₁ = g₁), + congr' with y x, + simp only [hf, hg] + end } @[simp] lemma coe_mk_apply (f₁ : α → option β) (f₂ : β → option α) (h) (x : α) : (pequiv.mk f₁ f₂ h : α → option β) x = f₁ x := rfl -@[ext] lemma ext : ∀ {f g : α ≃. β} (h : ∀ x, f x = g x), f = g -| ⟨f₁, f₂, hf⟩ ⟨g₁, g₂, hg⟩ h := -have h : f₁ = g₁, from funext h, -have ∀ b, f₂ b = g₂ b, - begin - subst h, - assume b, - have hf := λ a, hf a b, - have hg := λ a, hg a b, - cases h : g₂ b with a, - { simp only [h, option.not_mem_none, false_iff] at hg, - simp only [hg, iff_false] at hf, - rwa [option.eq_none_iff_forall_not_mem] }, - { rw [← option.mem_def, hf, ← hg, h, option.mem_def] } - end, -by simp [*, funext_iff] +@[ext] lemma ext {f g : α ≃. β} (h : ∀ x, f x = g x) : f = g := +fun_like.ext f g h -lemma ext_iff {f g : α ≃. β} : f = g ↔ ∀ x, f x = g x := -⟨congr_fun ∘ congr_arg _, ext⟩ +lemma ext_iff {f g : α ≃. β} : f = g ↔ ∀ x, f x = g x := fun_like.ext_iff /-- The identity map as a partial equivalence. -/ @[refl] protected def refl (α : Type*) : α ≃. α := diff --git a/src/order/initial_seg.lean b/src/order/initial_seg.lean index 16841b0aeddbe..4e3231e1c7e4d 100644 --- a/src/order/initial_seg.lean +++ b/src/order/initial_seg.lean @@ -56,7 +56,18 @@ localized "infix (name := initial_seg) ` ≼i `:25 := initial_seg" in initial_se namespace initial_seg instance : has_coe (r ≼i s) (r ↪r s) := ⟨initial_seg.to_rel_embedding⟩ -instance : has_coe_to_fun (r ≼i s) (λ _, α → β) := ⟨λ f x, (f : r ↪r s) x⟩ + +instance : embedding_like (r ≼i s) α β := +{ coe := λ f, f.to_fun, + coe_injective' := + begin + rintro ⟨f, hf⟩ ⟨g, hg⟩ h, + congr' with x, + exact congr_fun h x + end, + injective' := λ f, f.inj' } + +@[ext] lemma ext {f g : r ≼i s} (h : ∀ x, f x = g x) : f = g := fun_like.ext f g h @[simp] theorem coe_fn_mk (f : r ↪r s) (o) : (@initial_seg.mk _ _ r s f o : α → β) = f := rfl @@ -68,9 +79,11 @@ instance : has_coe_to_fun (r ≼i s) (λ _, α → β) := ⟨λ f x, (f : r ↪r theorem init' (f : r ≼i s) {a : α} {b : β} : s b (f a) → ∃ a', f a' = b := f.init _ _ +theorem map_rel_iff (f : r ≼i s) {a b : α} : s (f a) (f b) ↔ r a b := f.1.map_rel_iff + theorem init_iff (f : r ≼i s) {a : α} {b : β} : s b (f a) ↔ ∃ a', f a' = b ∧ r a' a := -⟨λ h, let ⟨a', e⟩ := f.init' h in ⟨a', e, (f : r ↪r s).map_rel_iff.1 (e.symm ▸ h)⟩, - λ ⟨a', e, h⟩, e ▸ (f : r ↪r s).map_rel_iff.2 h⟩ +⟨λ h, let ⟨a', e⟩ := f.init' h in ⟨a', e, f.map_rel_iff.1 (e.symm ▸ h)⟩, + λ ⟨a', e, h⟩, e ▸ f.map_rel_iff.2 h⟩ /-- An order isomorphism is an initial segment -/ def of_iso (f : r ≃r s) : r ≼i s := @@ -86,7 +99,7 @@ instance (r : α → α → Prop) : inhabited (r ≼i r) := ⟨initial_seg.refl @[trans] protected def trans (f : r ≼i s) (g : s ≼i t) : r ≼i t := ⟨f.1.trans g.1, λ a c h, begin simp at h ⊢, - rcases g.2 _ _ h with ⟨b, rfl⟩, have h := g.1.map_rel_iff.1 h, + rcases g.2 _ _ h with ⟨b, rfl⟩, have h := g.map_rel_iff.1 h, rcases f.2 _ _ h with ⟨a', rfl⟩, exact ⟨a', rfl⟩ end⟩ @@ -97,14 +110,11 @@ end⟩ theorem unique_of_trichotomous_of_irrefl [is_trichotomous β s] [is_irrefl β s] : well_founded r → subsingleton (r ≼i s) | ⟨h⟩ := ⟨λ f g, begin - suffices : (f : α → β) = g, { cases f, cases g, - congr, exact rel_embedding.coe_fn_injective this }, - funext a, have := h a, induction this with a H IH, - refine extensional_of_trichotomous_of_irrefl s (λ x, ⟨λ h, _, λ h, _⟩), - { rcases f.init_iff.1 h with ⟨y, rfl, h'⟩, - rw IH _ h', exact (g : r ↪r s).map_rel_iff.2 h' }, - { rcases g.init_iff.1 h with ⟨y, rfl, h'⟩, - rw ← IH _ h', exact (f : r ↪r s).map_rel_iff.2 h' } + ext a, + have := h a, induction this with a H IH, + refine extensional_of_trichotomous_of_irrefl s (λ x, _), + simp only [f.init_iff, g.init_iff], + exact exists_congr (λ x, and_congr_left $ λ hx, IH _ hx ▸ iff.rfl) end⟩ instance [is_well_order β s] : subsingleton (r ≼i s) := From 926daa81fd8acb2a04e15572c4ff20af2753c2ae Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Wed, 25 Jan 2023 11:06:55 +0000 Subject: [PATCH 0351/1291] feat(ring_theory/dedekind_domain): localizing a Dedekind domain at a prime gives DVR (#18264) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR shows one direction of the implication `is_dedekind_domain → is_dedekind_domain_dvr`. The first step is to show that any localization of a Dedekind domain is again Dedekind, then use the very useful `discrete_valuation_ring.tfae` to show that this implies the localization at a prime is a DVR. Beyond general usefulness, I need this for `ideal.span_norm`. Co-authored-by: Vierkantor Co-authored-by: Anne Baanen --- src/ring_theory/dedekind_domain/basic.lean | 11 +++ src/ring_theory/dedekind_domain/dvr.lean | 104 ++++++++++++++++++++- src/ring_theory/ideal/operations.lean | 3 + src/ring_theory/localization/ideal.lean | 12 +++ 4 files changed, 128 insertions(+), 2 deletions(-) diff --git a/src/ring_theory/dedekind_domain/basic.lean b/src/ring_theory/dedekind_domain/basic.lean index b5c2442176aab..c4ad4e9238297 100644 --- a/src/ring_theory/dedekind_domain/basic.lean +++ b/src/ring_theory/dedekind_domain/basic.lean @@ -65,6 +65,17 @@ lemma dimension_le_one.integral_closure [nontrivial R] [is_domain A] [algebra R (h : dimension_le_one R) : dimension_le_one (integral_closure R A) := h.is_integral_closure R A (integral_closure R A) +variables {R} + +lemma dimension_le_one.not_lt_lt (h : ring.dimension_le_one R) + (p₀ p₁ p₂ : ideal R) [hp₁ : p₁.is_prime] [hp₂ : p₂.is_prime] : + ¬ (p₀ < p₁ ∧ p₁ < p₂) +| ⟨h01, h12⟩ := h12.ne ((h p₁ (bot_le.trans_lt h01).ne' hp₁).eq_of_le hp₂.ne_top h12.le) + +lemma dimension_le_one.eq_bot_of_lt (h : ring.dimension_le_one R) + (p P : ideal R) [hp : p.is_prime] [hP : P.is_prime] (hpP : p < P) : p = ⊥ := +by_contra (λ hp0, h.not_lt_lt ⊥ p P ⟨ne.bot_lt hp0, hpP⟩) + end ring variables [is_domain A] diff --git a/src/ring_theory/dedekind_domain/dvr.lean b/src/ring_theory/dedekind_domain/dvr.lean index 0943227432454..4d4397ea2e091 100644 --- a/src/ring_theory/dedekind_domain/dvr.lean +++ b/src/ring_theory/dedekind_domain/dvr.lean @@ -3,21 +3,30 @@ 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 ring_theory.localization.at_prime +import ring_theory.dedekind_domain.ideal import ring_theory.discrete_valuation_ring +import ring_theory.localization.at_prime +import ring_theory.localization.submodule +import ring_theory.valuation.tfae /-! # Dedekind domains This file defines an equivalent notion of a Dedekind domain (or Dedekind ring), namely a Noetherian integral domain where the localization at all nonzero prime ideals is a DVR -(TODO: and shows that it is equivalent to the main definition). +(TODO: and shows that implies the main definition). ## Main definitions - `is_dedekind_domain_dvr` alternatively defines a Dedekind domain as an integral domain that is Noetherian, and the localization at every nonzero prime ideal is a DVR. +## Main results + - `is_localization.at_prime.discrete_valuation_ring_of_dedekind_domain` shows that + `is_dedekind_domain` implies the localization at each nonzero prime ideal is a DVR. + - `is_dedekind_domain.is_dedekind_domain_dvr` is one direction of the equivalence of definitions + of a Dedekind domain + ## Implementation notes The definitions that involve a field of fractions choose a canonical field of fractions, @@ -52,3 +61,94 @@ structure is_dedekind_domain_dvr : Prop := (is_noetherian_ring : is_noetherian_ring A) (is_dvr_at_nonzero_prime : ∀ P ≠ (⊥ : ideal A), P.is_prime → discrete_valuation_ring (localization.at_prime P)) + +/-- Localizing a domain of Krull dimension `≤ 1` gives another ring of Krull dimension `≤ 1`. + +Note that the same proof can/should be generalized to preserving any Krull dimension, +once we have a suitable definition. +-/ +lemma ring.dimension_le_one.localization {R : Type*} (Rₘ : Type*) [comm_ring R] [is_domain R] + [comm_ring Rₘ] [algebra R Rₘ] {M : submonoid R} [is_localization M Rₘ] (hM : M ≤ R⁰) + (h : ring.dimension_le_one R) : ring.dimension_le_one Rₘ := +begin + introsI p hp0 hpp, + refine ideal.is_maximal_def.mpr ⟨hpp.ne_top, ideal.maximal_of_no_maximal (λ P hpP hPm, _)⟩, + have hpP' : (⟨p, hpp⟩ : {p : ideal Rₘ // p.is_prime}) < ⟨P, hPm.is_prime⟩ := hpP, + rw ← (is_localization.order_iso_of_prime M Rₘ).lt_iff_lt at hpP', + haveI : ideal.is_prime (ideal.comap (algebra_map R Rₘ) p) := + ((is_localization.order_iso_of_prime M Rₘ) ⟨p, hpp⟩).2.1, + haveI : ideal.is_prime (ideal.comap (algebra_map R Rₘ) P) := + ((is_localization.order_iso_of_prime M Rₘ) ⟨P, hPm.is_prime⟩).2.1, + have hlt : ideal.comap (algebra_map R Rₘ) p < ideal.comap (algebra_map R Rₘ) P := hpP', + refine h.not_lt_lt ⊥ (ideal.comap _ _) (ideal.comap _ _) ⟨_, hpP'⟩, + exact is_localization.bot_lt_comap_prime _ _ hM _ hp0 +end + +/-- The localization of a Dedekind domain is a Dedekind domain. -/ +lemma is_localization.is_dedekind_domain [is_dedekind_domain A] {M : submonoid A} (hM : M ≤ A⁰) + (Aₘ : Type*) [comm_ring Aₘ] [is_domain Aₘ] [algebra A Aₘ] + [is_localization M Aₘ] : is_dedekind_domain Aₘ := +begin + have : ∀ (y : M), is_unit (algebra_map A (fraction_ring A) y), + { rintros ⟨y, hy⟩, + exact is_unit.mk0 _ (mt is_fraction_ring.to_map_eq_zero_iff.mp (non_zero_divisors.ne_zero + (hM hy))) }, + letI : algebra Aₘ (fraction_ring A) := ring_hom.to_algebra (is_localization.lift this), + haveI : is_scalar_tower A Aₘ (fraction_ring A) := is_scalar_tower.of_algebra_map_eq + (λ x, (is_localization.lift_eq this x).symm), + haveI : is_fraction_ring Aₘ (fraction_ring A) := + is_fraction_ring.is_fraction_ring_of_is_domain_of_is_localization M _ _, + refine (is_dedekind_domain_iff _ (fraction_ring A)).mpr ⟨_, _, _⟩, + { exact is_localization.is_noetherian_ring M _ (by apply_instance) }, + { exact is_dedekind_domain.dimension_le_one.localization Aₘ hM }, + { intros x hx, + obtain ⟨⟨y, y_mem⟩, hy⟩ := hx.exists_multiple_integral_of_is_localization M _, + obtain ⟨z, hz⟩ := (is_integrally_closed_iff _).mp is_dedekind_domain.is_integrally_closed hy, + refine ⟨is_localization.mk' Aₘ z ⟨y, y_mem⟩, (is_localization.lift_mk'_spec _ _ _ _).mpr _⟩, + rw [hz, set_like.coe_mk, ← algebra.smul_def], + refl }, +end + +/-- The localization of a Dedekind domain at every nonzero prime ideal is a Dedekind domain. -/ +lemma is_localization.at_prime.is_dedekind_domain [is_dedekind_domain A] + (P : ideal A) [P.is_prime] (Aₘ : Type*) [comm_ring Aₘ] [is_domain Aₘ] [algebra A Aₘ] + [is_localization.at_prime Aₘ P] : is_dedekind_domain Aₘ := +is_localization.is_dedekind_domain A P.prime_compl_le_non_zero_divisors Aₘ + +lemma is_localization.at_prime.not_is_field + {P : ideal A} (hP : P ≠ ⊥) [pP : P.is_prime] + (Aₘ : Type*) [comm_ring Aₘ] [algebra A Aₘ] [is_localization.at_prime Aₘ P] : + ¬ (is_field Aₘ) := +begin + intro h, + letI := h.to_field, + obtain ⟨x, x_mem, x_ne⟩ := P.ne_bot_iff.mp hP, + exact (local_ring.maximal_ideal.is_maximal _).ne_top (ideal.eq_top_of_is_unit_mem _ + ((is_localization.at_prime.to_map_mem_maximal_iff Aₘ P _).mpr x_mem) + (is_unit_iff_ne_zero.mpr ((map_ne_zero_iff (algebra_map A Aₘ) + (is_localization.injective Aₘ P.prime_compl_le_non_zero_divisors)).mpr x_ne))), +end + +/-- In a Dedekind domain, the localization at every nonzero prime ideal is a DVR. -/ +lemma is_localization.at_prime.discrete_valuation_ring_of_dedekind_domain [is_dedekind_domain A] + {P : ideal A} (hP : P ≠ ⊥) [pP : P.is_prime] + (Aₘ : Type*) [comm_ring Aₘ] [is_domain Aₘ] [algebra A Aₘ] [is_localization.at_prime Aₘ P] : + discrete_valuation_ring Aₘ := +begin + classical, + letI : is_noetherian_ring Aₘ := is_localization.is_noetherian_ring P.prime_compl _ + is_dedekind_domain.is_noetherian_ring, + letI : local_ring Aₘ := is_localization.at_prime.local_ring Aₘ P, + have hnf := is_localization.at_prime.not_is_field A hP Aₘ, + exact ((discrete_valuation_ring.tfae Aₘ hnf).out 0 2).mpr + (is_localization.at_prime.is_dedekind_domain A P _) +end + +/-- Dedekind domains, in the sense of Noetherian integrally closed domains of Krull dimension ≤ 1, +are also Dedekind domains in the sense of Noetherian domains where the localization at every +nonzero prime ideal is a DVR. -/ +theorem is_dedekind_domain.is_dedekind_domain_dvr [is_dedekind_domain A] : + is_dedekind_domain_dvr A := +{ is_noetherian_ring := is_dedekind_domain.is_noetherian_ring, + is_dvr_at_nonzero_prime := λ P hP pP, by exactI + is_localization.at_prime.discrete_valuation_ring_of_dedekind_domain A hP _ } diff --git a/src/ring_theory/ideal/operations.lean b/src/ring_theory/ideal/operations.lean index efde096489909..1d2ea816a284b 100644 --- a/src/ring_theory/ideal/operations.lean +++ b/src/ring_theory/ideal/operations.lean @@ -1309,6 +1309,9 @@ begin exact eq.symm (hf hx) ▸ (submodule.zero_mem ⊥) end +lemma comap_bot_of_injective : ideal.comap f ⊥ = ⊥ := +le_bot_iff.mp (ideal.comap_bot_le_of_injective f hf) + end injective end semiring diff --git a/src/ring_theory/localization/ideal.lean b/src/ring_theory/localization/ideal.lean index 2fdb7a3b9f348..4c07587b4a34f 100644 --- a/src/ring_theory/localization/ideal.lean +++ b/src/ring_theory/localization/ideal.lean @@ -190,6 +190,18 @@ begin (by rw [← ring_hom.map_mul, ← mk'_eq_mul_mk'_one, mk'_self, ring_hom.map_one]))) } end +open_locale non_zero_divisors + +lemma bot_lt_comap_prime [is_domain R] (hM : M ≤ R⁰) + (p : ideal S) [hpp : p.is_prime] (hp0 : p ≠ ⊥) : + ⊥ < ideal.comap (algebra_map R S) p := +begin + haveI : is_domain S := is_domain_of_le_non_zero_divisors _ hM, + convert (order_iso_of_prime M S).lt_iff_lt.mpr + (show (⟨⊥, ideal.bot_prime⟩ : {p : ideal S // p.is_prime}) < ⟨p, hpp⟩, from hp0.bot_lt), + exact (ideal.comap_bot_of_injective (algebra_map R S) (is_localization.injective _ hM)).symm, +end + end comm_ring end is_localization From 9ac58f083995faa064ac4d32c32264a4f7154825 Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Wed, 25 Jan 2023 11:06:57 +0000 Subject: [PATCH 0352/1291] feat(measure_theory/integrals): better interval_integral_pos (#18278) Currently `interval_integral_pos_of_pos` assumes the integrand is positive everywhere. This adds a version only assuming positivity on the domain of integration. --- .../integral/interval_integral.lean | 30 +++++++++++++------ 1 file changed, 21 insertions(+), 9 deletions(-) diff --git a/src/measure_theory/integral/interval_integral.lean b/src/measure_theory/integral/interval_integral.lean index 925d890f17376..8a99942ce3321 100644 --- a/src/measure_theory/integral/interval_integral.lean +++ b/src/measure_theory/integral/interval_integral.lean @@ -1260,18 +1260,30 @@ lemma integral_pos_iff_support_of_nonneg_ae (hf : 0 ≤ᵐ[μ] f) (hfi : interva 0 < ∫ x in a..b, f x ∂μ ↔ a < b ∧ 0 < μ (support f ∩ Ioc a b) := integral_pos_iff_support_of_nonneg_ae' (ae_mono measure.restrict_le_self hf) hfi -/-- If `f : ℝ → ℝ` is strictly positive and integrable on `(a, b]` for real numbers `a < b`, then -its integral over `a..b` is strictly positive. -/ -lemma interval_integral_pos_of_pos {f : ℝ → ℝ} {a b : ℝ} - (hfi : interval_integrable f measure_space.volume a b) (h : ∀ x, 0 < f x) (hab : a < b) : - 0 < ∫ x in a..b, f x := +/-- If `f : ℝ → ℝ` is integrable on `(a, b]` for real numbers `a < b`, and positive on the interior +of the interval, then its integral over `a..b` is strictly positive. -/ +lemma interval_integral_pos_of_pos_on {f : ℝ → ℝ} {a b : ℝ} + (hfi : interval_integrable f volume a b) (hpos : ∀ (x : ℝ), x ∈ Ioo a b → 0 < f x) (hab : a < b) : + 0 < ∫ (x : ℝ) in a..b, f x := begin - have hsupp : support f = univ := eq_univ_iff_forall.mpr (λ t, (h t).ne.symm), - replace h₀ : 0 ≤ᵐ[volume] f := eventually_of_forall (λ x, (h x).le), - rw integral_pos_iff_support_of_nonneg_ae h₀ hfi, - exact ⟨hab, by simp [hsupp, hab]⟩, + have hsupp : Ioo a b ⊆ support f ∩ Ioc a b := + λ x hx, ⟨mem_support.mpr (hpos x hx).ne', Ioo_subset_Ioc_self hx⟩, + have h₀ : 0 ≤ᵐ[volume.restrict (uIoc a b)] f, + { rw [eventually_le, uIoc_of_le hab.le], + refine ae_restrict_of_ae_eq_of_ae_restrict Ioo_ae_eq_Ioc _, + exact (ae_restrict_iff' measurable_set_Ioo).mpr (ae_of_all _ (λ x hx, (hpos x hx).le)) }, + rw integral_pos_iff_support_of_nonneg_ae' h₀ hfi, + exact ⟨hab, ((measure.measure_Ioo_pos _).mpr hab).trans_le (measure_mono hsupp)⟩, end +/-- If `f : ℝ → ℝ` is strictly positive everywhere, and integrable on `(a, b]` for real numbers +`a < b`, then its integral over `a..b` is strictly positive. (See `interval_integral_pos_of_pos_on` +for a version only assuming positivity of `f` on `(a, b)` rather than everywhere.) -/ +lemma interval_integral_pos_of_pos {f : ℝ → ℝ} {a b : ℝ} + (hfi : interval_integrable f measure_space.volume a b) (hpos : ∀ x, 0 < f x) (hab : a < b) : + 0 < ∫ x in a..b, f x := +interval_integral_pos_of_pos_on hfi (λ x hx, hpos x) hab + /-- If `f` and `g` are two functions that are interval integrable on `a..b`, `a ≤ b`, `f x ≤ g x` for a.e. `x ∈ set.Ioc a b`, and `f x < g x` on a subset of `set.Ioc a b` of nonzero measure, then `∫ x in a..b, f x ∂μ < ∫ x in a..b, g x ∂μ`. -/ From 4d75632823d091877e5e63cc5694c8cb807cb7d4 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Wed, 25 Jan 2023 14:10:22 +0000 Subject: [PATCH 0353/1291] feat(order/game_add): add more lemmas (#16082) This PR does the following - add miscellaneous lemmas on `prod.game_add` - add a two-variable recursion principle for `prod.game_add` Mathlib4 pair: https://github.com/leanprover-community/mathlib4/pull/1825 --- src/order/game_add.lean | 51 +++++++++++++++++++++++++++++++++++++---- 1 file changed, 47 insertions(+), 4 deletions(-) diff --git a/src/order/game_add.lean b/src/order/game_add.lean index 0a9f517cd7ade..d8e8a15a9f702 100644 --- a/src/order/game_add.lean +++ b/src/order/game_add.lean @@ -25,7 +25,6 @@ subsequency relation on the addition of combinatorial games. ## Todo -- Add custom `induction` and `fix` lemmas. - Define `sym2.game_add`. -/ @@ -45,12 +44,31 @@ inductive game_add : α × β → α × β → Prop | fst {a₁ a₂ b} : rα a₁ a₂ → game_add (a₁, b) (a₂, b) | snd {a b₁ b₂} : rβ b₁ b₂ → game_add (a, b₁) (a, b₂) -/-- `game_add` is a `subrelation` of `prod.lex`. -/ +lemma game_add_iff {rα rβ} {x y : α × β} : + game_add rα rβ x y ↔ rα x.1 y.1 ∧ x.2 = y.2 ∨ rβ x.2 y.2 ∧ x.1 = y.1 := +begin + split, + { rintro (@⟨a₁, a₂, b, h⟩ | @⟨a, b₁, b₂, h⟩), + exacts [or.inl ⟨h, rfl⟩, or.inr ⟨h, rfl⟩] }, + { revert x y, + rintro ⟨a₁, b₁⟩ ⟨a₂, b₂⟩ (⟨h, rfl : b₁ = b₂⟩ | ⟨h, rfl : a₁ = a₂⟩), + exacts [game_add.fst h, game_add.snd h] } +end + +lemma game_add_mk_iff {rα rβ} {a₁ a₂ : α} {b₁ b₂ : β} : + game_add rα rβ (a₁, b₁) (a₂, b₂) ↔ rα a₁ a₂ ∧ b₁ = b₂ ∨ rβ b₁ b₂ ∧ a₁ = a₂ := +game_add_iff + +@[simp] lemma game_add_swap_swap : ∀ a b : α × β, + game_add rβ rα a.swap b.swap ↔ game_add rα rβ a b := +λ ⟨a₁, b₁⟩ ⟨a₂, b₂⟩, by rw [prod.swap, game_add_mk_iff, game_add_mk_iff, or_comm] + +/-- `prod.game_add` is a `subrelation` of `prod.lex`. -/ lemma game_add_le_lex : game_add rα rβ ≤ prod.lex rα rβ := λ _ _ h, h.rec (λ _ _ b, prod.lex.left b b) (λ a _ _, prod.lex.right a) -/-- `prod.rprod` is a subrelation of the transitive closure of `game_add`. -/ -lemma rprod_le_trans_gen_game_add : prod.rprod rα rβ ≤ relation.trans_gen (game_add rα rβ) := +/-- `prod.rprod` is a subrelation of the transitive closure of `prod.game_add`. -/ +lemma rprod_le_trans_gen_game_add : rprod rα rβ ≤ relation.trans_gen (game_add rα rβ) := λ _ _ h, h.rec begin intros _ _ _ _ hα hβ, exact relation.trans_gen.tail (relation.trans_gen.single $ game_add.fst hα) (game_add.snd hβ), @@ -77,3 +95,28 @@ end In particular, the sum of two well-founded games is well-founded. -/ lemma well_founded.prod_game_add (hα : well_founded rα) (hβ : well_founded rβ) : well_founded (prod.game_add rα rβ) := ⟨λ ⟨a, b⟩, (hα.apply a).prod_game_add (hβ.apply b)⟩ + +namespace prod + +/-- Recursion on the well-founded `prod.game_add` relation. + + Note that it's strictly more general to recurse on the lexicographic order instead. -/ +def game_add.fix {C : α → β → Sort*} (hα : well_founded rα) (hβ : well_founded rβ) + (IH : Π a₁ b₁, (Π a₂ b₂, game_add rα rβ (a₂, b₂) (a₁, b₁) → C a₂ b₂) → C a₁ b₁) (a : α) (b : β) : + C a b := +@well_founded.fix (α × β) (λ x, C x.1 x.2) _ (hα.prod_game_add hβ) + (λ ⟨x₁, x₂⟩ IH', IH x₁ x₂ $ λ a' b', IH' ⟨a', b'⟩) ⟨a, b⟩ + +lemma game_add.fix_eq {C : α → β → Sort*} (hα : well_founded rα) (hβ : well_founded rβ) + (IH : Π a₁ b₁, (Π a₂ b₂, game_add rα rβ (a₂, b₂) (a₁, b₁) → C a₂ b₂) → C a₁ b₁) (a : α) (b : β) : + game_add.fix hα hβ IH a b = IH a b (λ a' b' h, game_add.fix hα hβ IH a' b') := +by { rw [game_add.fix, well_founded.fix_eq], refl } + +/-- Induction on the well-founded `prod.game_add` relation. + + Note that it's strictly more general to induct on the lexicographic order instead. -/ +lemma game_add.induction {C : α → β → Prop} : well_founded rα → well_founded rβ → + (∀ a₁ b₁, (∀ a₂ b₂, game_add rα rβ (a₂, b₂) (a₁, b₁) → C a₂ b₂) → C a₁ b₁) → ∀ a b, C a b := +game_add.fix + +end prod From a239cd3e7ac2c7cde36c913808f9d40c411344f6 Mon Sep 17 00:00:00 2001 From: Siddhartha Prasad Date: Wed, 25 Jan 2023 14:10:24 +0000 Subject: [PATCH 0354/1291] feat(algebra/order/kleene) : Kleene algebras (#17965) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Define idempotent semirings and Kleene algebras, which are used extensively in the theory of computation. Co-authored-by: Siddhartha Prasad Co-authored-by: Siddhartha Prasad Co-authored-by: Siddhartha Prasad Co-authored-by: Yaël Dillies --- docs/references.bib | 16 ++ src/algebra/order/kleene.lean | 291 +++++++++++++++++++++ src/computability/DFA.lean | 9 +- src/computability/NFA.lean | 3 +- src/computability/epsilon_NFA.lean | 4 +- src/computability/language.lean | 92 +++---- src/computability/regular_expressions.lean | 9 +- 7 files changed, 369 insertions(+), 55 deletions(-) create mode 100644 src/algebra/order/kleene.lean diff --git a/docs/references.bib b/docs/references.bib index 32680882acc87..3d2220c570725 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -1405,6 +1405,22 @@ @Article{ KoukoulopoulosMaynard2020 url = {https://doi.org/10.4007/annals.2020.192.1.5}, } +@Article{ kozen1994, + title = {A Completeness Theorem for Kleene Algebras and the Algebra of Regular Events}, + journal = {Information and Computation}, + volume = {110}, + number = {2}, + pages = {366-390}, + year = {1994}, + issn = {0890-5401}, + doi = {https://doi.org/10.1006/inco.1994.1037}, + url = {https://www.sciencedirect.com/science/article/pii/S0890540184710376}, + author = {D. Kozen}, + abstract = {We give a finitary axiomatization of the algebra of regular events involving only + equations and equational implications. Unlike Salomaa′s axiomatizations, the + axiomatization given here is sound for all interpretations over Kleene algebras.} +} + @Article{ lazarus1973, author = {Michel Lazarus}, title = {Les familles libres maximales d'un module ont-elles le diff --git a/src/algebra/order/kleene.lean b/src/algebra/order/kleene.lean new file mode 100644 index 0000000000000..0a2487a126cdf --- /dev/null +++ b/src/algebra/order/kleene.lean @@ -0,0 +1,291 @@ +/- +Copyright (c) 2022 Siddhartha Prasad, Yaël Dillies. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Siddhartha Prasad, Yaël Dillies +-/ +import algebra.order.ring.canonical +import algebra.ring.pi +import algebra.ring.prod +import order.hom.complete_lattice + +/-! +# Kleene Algebras + +This file defines idempotent semirings and Kleene algebras, which are used extensively in the theory +of computation. + +An idempotent semiring is a semiring whose addition is idempotent. An idempotent semiring is +naturally a semilattice by setting `a ≤ b` if `a + b = b`. + +A Kleene algebra is an idempotent semiring equipped with an additional unary operator `∗`, the +Kleene star. + +## Main declarations + +* `idem_semiring`: Idempotent semiring +* `idem_comm_semiring`: Idempotent commutative semiring +* `kleene_algebra`: Kleene algebra + +## Notation + +`a∗` is notation for `kstar a` in locale `computability`. + +## References + +* [D. Kozen, *A completeness theorem for Kleene algebras and the algebra of regular events*] + [kozen1994] +* https://planetmath.org/idempotentsemiring +* https://encyclopediaofmath.org/wiki/Idempotent_semi-ring +* https://planetmath.org/kleene_algebra + +## TODO + +Instances for `add_opposite`, `mul_opposite`, `ulift`, `subsemiring`, `subring`, `subalgebra`. + +## Tags + +kleene algebra, idempotent semiring +-/ + +set_option old_structure_cmd true + +open function + +universe u +variables {α β ι : Type*} {π : ι → Type*} + +/-- An idempotent semiring is a semiring with the additional property that addition is idempotent. +-/ +@[protect_proj] +class idem_semiring (α : Type u) extends semiring α, semilattice_sup α := +(sup := (+)) +(add_eq_sup : ∀ a b : α, a + b = a ⊔ b . try_refl_tac) +(bot : α := 0) +(bot_le : ∀ a, bot ≤ a) + +/-- An idempotent commutative semiring is a commutative semiring with the additional property that +addition is idempotent. -/ +@[protect_proj] +class idem_comm_semiring (α : Type u) extends comm_semiring α, idem_semiring α + +/-- Notation typeclass for the Kleene star `∗`. -/ +@[protect_proj] +class has_kstar (α : Type*) := +(kstar : α → α) + +localized "postfix `∗`:1025 := has_kstar.kstar" in computability + +/-- A Kleene Algebra is an idempotent semiring with an additional unary operator `kstar` (for Kleene +star) that satisfies the following properties: +* `1 + a * a∗ ≤ a∗` +* `1 + a∗ * a ≤ a∗` +* If `a * c + b ≤ c`, then `a∗ * b ≤ c` +* If `c * a + b ≤ c`, then `b * a∗ ≤ c` +-/ +@[protect_proj] +class kleene_algebra (α : Type*) extends idem_semiring α, has_kstar α := +(one_le_kstar : ∀ a : α, 1 ≤ a∗) +(mul_kstar_le_kstar : ∀ a : α, a * a∗ ≤ a∗) +(kstar_mul_le_kstar : ∀ a : α, a∗ * a ≤ a∗) +(mul_kstar_le_self : ∀ a b : α, b * a ≤ b → b * a∗ ≤ b) +(kstar_mul_le_self : ∀ a b : α, a * b ≤ b → a∗ * b ≤ b) + +@[priority 100] -- See note [lower instance priority] +instance idem_semiring.to_order_bot [idem_semiring α] : order_bot α := { ..‹idem_semiring α› } + +/-- Construct an idempotent semiring from an idempotent addition. -/ +@[reducible] -- See note [reducible non-instances] +def idem_semiring.of_semiring [semiring α] (h : ∀ a : α, a + a = a) : idem_semiring α := +{ le := λ a b, a + b = b, + le_refl := h, + le_trans := λ a b c (hab : _ = _) (hbc : _ = _), by { change _ = _, rw [←hbc, ←add_assoc, hab] }, + le_antisymm := λ a b (hab : _ = _) (hba : _ = _), by rwa [←hba, add_comm], + sup := (+), + le_sup_left := λ a b, by { change _ = _, rw [←add_assoc, h] }, + le_sup_right := λ a b, by { change _ = _, rw [add_comm, add_assoc, h] }, + sup_le := λ a b c hab (hbc : _ = _), by { change _ = _, rwa [add_assoc, hbc] }, + bot := 0, + bot_le := zero_add, + ..‹semiring α› } + +section idem_semiring +variables [idem_semiring α] {a b c : α} + +@[simp] lemma add_eq_sup (a b : α) : a + b = a ⊔ b := idem_semiring.add_eq_sup _ _ +lemma add_idem (a : α) : a + a = a := by simp + +lemma nsmul_eq_self : ∀ {n : ℕ} (hn : n ≠ 0) (a : α), n • a = a +| 0 h := (h rfl).elim +| 1 h := one_nsmul +| (n + 2) h := λ a, by rw [succ_nsmul, nsmul_eq_self n.succ_ne_zero, add_idem] + +lemma add_eq_left_iff_le : a + b = a ↔ b ≤ a := by simp +lemma add_eq_right_iff_le : a + b = b ↔ a ≤ b := by simp + +alias add_eq_left_iff_le ↔ _ has_le.le.add_eq_left +alias add_eq_right_iff_le ↔ _ has_le.le.add_eq_right + +lemma add_le_iff : a + b ≤ c ↔ a ≤ c ∧ b ≤ c := by simp +lemma add_le (ha : a ≤ c) (hb : b ≤ c) : a + b ≤ c := add_le_iff.2 ⟨ha, hb⟩ + +@[priority 100] -- See note [lower instance priority] +instance idem_semiring.to_canonically_ordered_add_monoid : canonically_ordered_add_monoid α := +{ add_le_add_left := λ a b hbc c, by { simp_rw add_eq_sup, exact sup_le_sup_left hbc _ }, + exists_add_of_le := λ a b h, ⟨b, h.add_eq_right.symm⟩, + le_self_add := λ a b, add_eq_right_iff_le.1 $ by rw [←add_assoc, add_idem], + ..‹idem_semiring α› } + +@[priority 100] -- See note [lower instance priority] +instance idem_semiring.to_covariant_class_mul_le : covariant_class α α (*) (≤) := +⟨λ a b c hbc, add_eq_left_iff_le.1 $ by rw [←mul_add, hbc.add_eq_left]⟩ + +@[priority 100] -- See note [lower instance priority] +instance idem_semiring.to_covariant_class_swap_mul_le : covariant_class α α (swap (*)) (≤) := +⟨λ a b c hbc, add_eq_left_iff_le.1 $ by rw [←add_mul, hbc.add_eq_left]⟩ + +end idem_semiring + +section kleene_algebra +variables [kleene_algebra α] {a b c : α} + +@[simp] lemma one_le_kstar : 1 ≤ a∗ := kleene_algebra.one_le_kstar _ +lemma mul_kstar_le_kstar : a * a∗ ≤ a∗ := kleene_algebra.mul_kstar_le_kstar _ +lemma kstar_mul_le_kstar : a∗ * a ≤ a∗ := kleene_algebra.kstar_mul_le_kstar _ +lemma mul_kstar_le_self : b * a ≤ b → b * a∗ ≤ b := kleene_algebra.mul_kstar_le_self _ _ +lemma kstar_mul_le_self : a * b ≤ b → a∗ * b ≤ b := kleene_algebra.kstar_mul_le_self _ _ + +lemma mul_kstar_le (hb : b ≤ c) (ha : c * a ≤ c) : b * a∗ ≤ c := +(mul_le_mul_right' hb _).trans $ mul_kstar_le_self ha + +lemma kstar_mul_le (hb : b ≤ c) (ha : a * c ≤ c) : a∗ * b ≤ c := +(mul_le_mul_left' hb _).trans $ kstar_mul_le_self ha + +lemma kstar_le_of_mul_le_left (hb : 1 ≤ b) : b * a ≤ b → a∗ ≤ b := by simpa using mul_kstar_le hb +lemma kstar_le_of_mul_le_right (hb : 1 ≤ b) : a * b ≤ b → a∗ ≤ b := by simpa using kstar_mul_le hb + +@[simp] lemma le_kstar : a ≤ a∗ := le_trans (le_mul_of_one_le_left' one_le_kstar) kstar_mul_le_kstar + +@[mono] lemma kstar_mono : monotone (has_kstar.kstar : α → α) := +λ a b h, kstar_le_of_mul_le_left one_le_kstar $ kstar_mul_le (h.trans le_kstar) $ + mul_kstar_le_kstar + +@[simp] lemma kstar_eq_one : a∗ = 1 ↔ a ≤ 1 := +⟨le_kstar.trans_eq, λ h, one_le_kstar.antisymm' $ kstar_le_of_mul_le_left le_rfl $ by rwa one_mul⟩ + +@[simp] lemma kstar_zero : (0 : α)∗ = 1 := kstar_eq_one.2 zero_le_one +@[simp] lemma kstar_one : (1 : α)∗ = 1 := kstar_eq_one.2 le_rfl + +@[simp] lemma kstar_mul_kstar (a : α) : a∗ * a∗ = a∗ := +(mul_kstar_le le_rfl $ kstar_mul_le_kstar).antisymm $ le_mul_of_one_le_left' one_le_kstar + +@[simp] lemma kstar_eq_self : a∗ = a ↔ a * a = a ∧ 1 ≤ a := +⟨λ h, ⟨by rw [←h, kstar_mul_kstar], one_le_kstar.trans_eq h⟩, λ h, + (kstar_le_of_mul_le_left h.2 h.1.le).antisymm le_kstar⟩ + +@[simp] lemma kstar_idem (a : α) : a∗∗ = a∗ := kstar_eq_self.2 ⟨kstar_mul_kstar _, one_le_kstar⟩ + +@[simp] lemma pow_le_kstar : ∀ {n : ℕ}, a ^ n ≤ a∗ +| 0 := (pow_zero _).trans_le one_le_kstar +| (n + 1) := by {rw pow_succ, exact (mul_le_mul_left' pow_le_kstar _).trans mul_kstar_le_kstar } + +end kleene_algebra + +namespace prod + +instance [idem_semiring α] [idem_semiring β] : idem_semiring (α × β) := +{ add_eq_sup := λ a b, ext (add_eq_sup _ _) (add_eq_sup _ _), + ..prod.semiring, ..prod.semilattice_sup _ _, ..prod.order_bot _ _ } + +instance [idem_comm_semiring α] [idem_comm_semiring β] : idem_comm_semiring (α × β) := +{ ..prod.comm_semiring, ..prod.idem_semiring } + +variables [kleene_algebra α] [kleene_algebra β] + +instance : kleene_algebra (α × β) := +{ kstar := λ a, (a.1∗, a.2∗), + one_le_kstar := λ a, ⟨one_le_kstar, one_le_kstar⟩, + mul_kstar_le_kstar := λ a, ⟨mul_kstar_le_kstar, mul_kstar_le_kstar⟩, + kstar_mul_le_kstar := λ a, ⟨kstar_mul_le_kstar, kstar_mul_le_kstar⟩, + mul_kstar_le_self := λ a b, and.imp mul_kstar_le_self mul_kstar_le_self, + kstar_mul_le_self := λ a b, and.imp kstar_mul_le_self kstar_mul_le_self, + ..prod.idem_semiring } + +lemma kstar_def (a : α × β) : a∗ = (a.1∗, a.2∗) := rfl +@[simp] lemma fst_kstar (a : α × β) : a∗.1 = a.1∗ := rfl +@[simp] lemma snd_kstar (a : α × β) : a∗.2 = a.2∗ := rfl + +end prod + +namespace pi + +instance [Π i, idem_semiring (π i)] : idem_semiring (Π i, π i) := +{ add_eq_sup := λ a b, funext $ λ i, add_eq_sup _ _, + ..pi.semiring, ..pi.semilattice_sup, ..pi.order_bot } + +instance [Π i, idem_comm_semiring (π i)] : idem_comm_semiring (Π i, π i) := +{ ..pi.comm_semiring, ..pi.idem_semiring } + +variables [Π i, kleene_algebra (π i)] + +instance : kleene_algebra (Π i, π i) := +{ kstar := λ a i, (a i)∗, + one_le_kstar := λ a i, one_le_kstar, + mul_kstar_le_kstar := λ a i, mul_kstar_le_kstar, + kstar_mul_le_kstar := λ a i, kstar_mul_le_kstar, + mul_kstar_le_self := λ a b h i, mul_kstar_le_self $ h _, + kstar_mul_le_self := λ a b h i, kstar_mul_le_self $ h _, + ..pi.idem_semiring } + +lemma kstar_def (a : Π i, π i) : a∗ = λ i, (a i)∗ := rfl +@[simp] lemma kstar_apply (a : Π i, π i) (i : ι) : a∗ i = (a i)∗ := rfl + +end pi + +namespace function.injective + +/-- Pullback an `idem_semiring` instance along an injective function. -/ +@[reducible] -- See note [reducible non-instances] +protected def idem_semiring [idem_semiring α] [has_zero β] [has_one β] [has_add β] [has_mul β] + [has_pow β ℕ] [has_smul ℕ β] [has_nat_cast β] [has_sup β] [has_bot β] + (f : β → α) (hf : injective f) (zero : f 0 = 0) (one : f 1 = 1) + (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) + (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) + (nat_cast : ∀ n : ℕ, f n = n) (sup : ∀ a b, f (a ⊔ b) = f a ⊔ f b) (bot : f ⊥ = ⊥) : + idem_semiring β := +{ add_eq_sup := λ a b, hf $ by erw [sup, add, add_eq_sup], + bot := ⊥, + bot_le := λ a, bot.trans_le $ @bot_le _ _ _ $ f a, + ..hf.semiring f zero one add mul nsmul npow nat_cast, ..hf.semilattice_sup _ sup, ..‹has_bot β› } + +/-- Pullback an `idem_comm_semiring` instance along an injective function. -/ +@[reducible] -- See note [reducible non-instances] +protected def idem_comm_semiring [idem_comm_semiring α] [has_zero β] [has_one β] [has_add β] + [has_mul β] [has_pow β ℕ] [has_smul ℕ β] [has_nat_cast β] [has_sup β] [has_bot β] + (f : β → α) (hf : injective f) (zero : f 0 = 0) (one : f 1 = 1) + (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) + (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) + (nat_cast : ∀ n : ℕ, f n = n) (sup : ∀ a b, f (a ⊔ b) = f a ⊔ f b) (bot : f ⊥ = ⊥) : + idem_comm_semiring β := +{ ..hf.comm_semiring f zero one add mul nsmul npow nat_cast, + ..hf.idem_semiring f zero one add mul nsmul npow nat_cast sup bot } + +/-- Pullback an `idem_comm_semiring` instance along an injective function. -/ +@[reducible] -- See note [reducible non-instances] +protected def kleene_algebra [kleene_algebra α] [has_zero β] [has_one β] [has_add β] + [has_mul β] [has_pow β ℕ] [has_smul ℕ β] [has_nat_cast β] [has_sup β] [has_bot β] [has_kstar β] + (f : β → α) (hf : injective f) (zero : f 0 = 0) (one : f 1 = 1) + (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) + (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) + (nat_cast : ∀ n : ℕ, f n = n) (sup : ∀ a b, f (a ⊔ b) = f a ⊔ f b) (bot : f ⊥ = ⊥) + (kstar : ∀ a, f a∗ = (f a)∗) : + kleene_algebra β := +{ one_le_kstar := λ a, one.trans_le $ by { erw kstar, exact one_le_kstar }, + mul_kstar_le_kstar := λ a, by { change f _ ≤ _, erw [mul, kstar], exact mul_kstar_le_kstar }, + kstar_mul_le_kstar := λ a, by { change f _ ≤ _, erw [mul, kstar], exact kstar_mul_le_kstar }, + mul_kstar_le_self := λ a b (h : f _ ≤ _), + by { change f _ ≤ _, erw [mul, kstar], erw mul at h, exact mul_kstar_le_self h }, + kstar_mul_le_self := λ a b (h : f _ ≤ _), + by { change f _ ≤ _, erw [mul, kstar], erw mul at h, exact kstar_mul_le_self h }, + ..hf.idem_semiring f zero one add mul nsmul npow nat_cast sup bot, ..‹has_kstar β› } + +end function.injective diff --git a/src/computability/DFA.lean b/src/computability/DFA.lean index 57ab5cfe17b0e..0c815938b9a00 100644 --- a/src/computability/DFA.lean +++ b/src/computability/DFA.lean @@ -10,6 +10,7 @@ import tactic.norm_num /-! # Deterministic Finite Automata + This file contains the definition of a Deterministic Finite Automaton (DFA), a state machine which determines whether a string (implemented as a list over an arbitrary alphabet) is in a regular set in linear time. @@ -17,6 +18,8 @@ Note that this definition allows for Automaton with infinite states, a `fintype` supplied for true DFA's. -/ +open_locale computability + universes u v /-- A DFA is a set of states (`σ`), a transition function from state to state labelled by the @@ -109,9 +112,9 @@ begin end lemma eval_from_of_pow {x y : list α} {s : σ} (hx : M.eval_from s x = s) - (hy : y ∈ @language.star α {x}) : M.eval_from s y = s := + (hy : y ∈ ({x} : language α)∗) : M.eval_from s y = s := begin - rw language.mem_star at hy, + rw language.mem_kstar at hy, rcases hy with ⟨ S, rfl, hS ⟩, induction S with a S ih, { refl }, @@ -126,7 +129,7 @@ end lemma pumping_lemma [fintype σ] {x : list α} (hx : x ∈ M.accepts) (hlen : fintype.card σ ≤ list.length x) : ∃ a b c, x = a ++ b ++ c ∧ a.length + b.length ≤ fintype.card σ ∧ b ≠ [] ∧ - {a} * language.star {b} * {c} ≤ M.accepts := + {a} * {b}∗ * {c} ≤ M.accepts := begin obtain ⟨_, a, b, c, hx, hlen, hnil, rfl, hb, hc⟩ := M.eval_from_split hlen rfl, use [a, b, c, hx, hlen, hnil], diff --git a/src/computability/NFA.lean b/src/computability/NFA.lean index ab90f0f21d938..b1fce95851172 100644 --- a/src/computability/NFA.lean +++ b/src/computability/NFA.lean @@ -18,6 +18,7 @@ supplied for true NFA's. -/ open set +open_locale computability universes u v @@ -89,7 +90,7 @@ end lemma pumping_lemma [fintype σ] {x : list α} (hx : x ∈ M.accepts) (hlen : fintype.card (set σ) ≤ list.length x) : ∃ a b c, x = a ++ b ++ c ∧ a.length + b.length ≤ fintype.card (set σ) ∧ b ≠ [] ∧ - {a} * language.star {b} * {c} ≤ M.accepts := + {a} * {b}∗ * {c} ≤ M.accepts := begin rw ←to_DFA_correct at hx ⊢, exact M.to_DFA.pumping_lemma hx hlen diff --git a/src/computability/epsilon_NFA.lean b/src/computability/epsilon_NFA.lean index 1b81ac03f9390..6faa55eafe884 100644 --- a/src/computability/epsilon_NFA.lean +++ b/src/computability/epsilon_NFA.lean @@ -8,6 +8,7 @@ import computability.NFA /-! # Epsilon Nondeterministic Finite Automata + This file contains the definition of an epsilon Nondeterministic Finite Automaton (`ε_NFA`), a state machine which determines whether a string (implemented as a list over an arbitrary alphabet) is in a regular set by evaluating the string over every possible path, also having access to ε-transitons, @@ -17,6 +18,7 @@ supplied for true `ε_NFA`'s. -/ open set +open_locale computability universes u v @@ -116,7 +118,7 @@ end lemma pumping_lemma [fintype σ] {x : list α} (hx : x ∈ M.accepts) (hlen : fintype.card (set σ) ≤ list.length x) : ∃ a b c, x = a ++ b ++ c ∧ a.length + b.length ≤ fintype.card (set σ) ∧ b ≠ [] ∧ - {a} * language.star {b} * {c} ≤ M.accepts := + {a} * {b}∗ * {c} ≤ M.accepts := begin rw ←to_NFA_correct at hx ⊢, exact M.to_NFA.pumping_lemma hx hlen diff --git a/src/computability/language.lean b/src/computability/language.lean index 2a2f1e673acfd..ae2ef1a0e5ef4 100644 --- a/src/computability/language.lean +++ b/src/computability/language.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Fox Thomson -/ import algebra.hom.ring +import algebra.order.kleene import data.list.join import data.set.lattice @@ -17,6 +18,7 @@ over the languages. -/ open list set +open_locale computability universes v @@ -51,23 +53,23 @@ lemma one_def : (1 : language α) = {[]} := rfl lemma add_def (l m : language α) : l + m = l ∪ m := rfl lemma mul_def (l m : language α) : l * m = image2 (++) l m := rfl -/-- The star of a language `L` is the set of all strings which can be written by concatenating - strings from `L`. -/ -def star (l : language α) : language α := -{ x | ∃ S : list (list α), x = S.join ∧ ∀ y ∈ S, y ∈ l} +/-- The Kleene star of a language `L` is the set of all strings which can be written by +concatenating strings from `L`. -/ +instance : has_kstar (language α) := ⟨λ l, {x | ∃ L : list (list α), x = L.join ∧ ∀ y ∈ L, y ∈ l}⟩ + +lemma kstar_def (l : language α) : l∗ = {x | ∃ L : list (list α), x = L.join ∧ ∀ y ∈ L, y ∈ l} := +rfl -lemma star_def (l : language α) : - l.star = { x | ∃ S : list (list α), x = S.join ∧ ∀ y ∈ S, y ∈ l} := rfl @[simp] lemma not_mem_zero (x : list α) : x ∉ (0 : language α) := id @[simp] lemma mem_one (x : list α) : x ∈ (1 : language α) ↔ x = [] := by refl lemma nil_mem_one : [] ∈ (1 : language α) := set.mem_singleton _ -@[simp] lemma mem_add (l m : language α) (x : list α) : x ∈ l + m ↔ x ∈ l ∨ x ∈ m := iff.rfl +lemma mem_add (l m : language α) (x : list α) : x ∈ l + m ↔ x ∈ l ∨ x ∈ m := iff.rfl lemma mem_mul : x ∈ l * m ↔ ∃ a b, a ∈ l ∧ b ∈ m ∧ a ++ b = x := mem_image2 lemma append_mem_mul : a ∈ l → b ∈ m → a ++ b ∈ l * m := mem_image2_of_mem -lemma mem_star : x ∈ l.star ↔ ∃ S : list (list α), x = S.join ∧ ∀ y ∈ S, y ∈ l := iff.rfl -lemma join_mem_star {S : list (list α)} (h : ∀ y ∈ S, y ∈ l) : S.join ∈ l.star := ⟨S, rfl, h⟩ -lemma nil_mem_star (l : language α) : [] ∈ l.star := ⟨[], rfl, λ _, false.elim⟩ +lemma mem_kstar : x ∈ l∗ ↔ ∃ L : list (list α), x = L.join ∧ ∀ y ∈ L, y ∈ l := iff.rfl +lemma join_mem_kstar {L : list (list α)} (h : ∀ y ∈ L, y ∈ l) : L.join ∈ l∗ := ⟨L, rfl, h⟩ +lemma nil_mem_kstar (l : language α) : [] ∈ l∗ := ⟨[], rfl, λ _, false.elim⟩ instance : semiring (language α) := { add := (+), @@ -103,8 +105,8 @@ def map (f : α → β) : language α →+* language β := @[simp] lemma map_map (g : β → γ) (f : α → β) (l : language α) : map g (map f l) = map (g ∘ f) l := by simp [map, image_image] -lemma star_def_nonempty (l : language α) : - l.star = {x | ∃ S : list (list α), x = S.join ∧ ∀ y ∈ S, y ∈ l ∧ y ≠ []} := +lemma kstar_def_nonempty (l : language α) : + l∗ = {x | ∃ S : list (list α), x = S.join ∧ ∀ y ∈ S, y ∈ l ∧ y ≠ []} := begin ext x, split, @@ -162,56 +164,54 @@ begin exact ⟨a, _, hS.1, ⟨S, rfl, rfl, hS.2⟩, rfl⟩ } } end -lemma star_eq_supr_pow (l : language α) : l.star = ⨆ i : ℕ, l ^ i := +lemma kstar_eq_supr_pow (l : language α) : l∗ = ⨆ i : ℕ, l ^ i := begin ext x, - simp only [mem_star, mem_supr, mem_pow], + simp only [mem_kstar, mem_supr, mem_pow], split, { rintro ⟨S, rfl, hS⟩, exact ⟨_, S, rfl, rfl, hS⟩ }, { rintro ⟨_, S, rfl, rfl, hS⟩, exact ⟨S, rfl, hS⟩ } end -@[simp] lemma map_star (f : α → β) (l : language α) : map f (star l) = star (map f l) := +@[simp] lemma map_kstar (f : α → β) (l : language α) : map f l∗ = (map f l)∗ := begin - rw [star_eq_supr_pow, star_eq_supr_pow], + rw [kstar_eq_supr_pow, kstar_eq_supr_pow], simp_rw ←map_pow, exact image_Union, end -lemma mul_self_star_comm (l : language α) : l.star * l = l * l.star := -by simp only [star_eq_supr_pow, mul_supr, supr_mul, ← pow_succ, ← pow_succ'] +lemma mul_self_kstar_comm (l : language α) : l∗ * l = l * l∗ := +by simp only [kstar_eq_supr_pow, mul_supr, supr_mul, ← pow_succ, ← pow_succ'] -@[simp] lemma one_add_self_mul_star_eq_star (l : language α) : 1 + l * l.star = l.star := +@[simp] lemma one_add_self_mul_kstar_eq_kstar (l : language α) : 1 + l * l∗ = l∗ := begin - simp only [star_eq_supr_pow, mul_supr, ← pow_succ, ← pow_zero l], + simp only [kstar_eq_supr_pow, mul_supr, ← pow_succ, ← pow_zero l], exact sup_supr_nat_succ _ end -@[simp] lemma one_add_star_mul_self_eq_star (l : language α) : 1 + l.star * l = l.star := -by rw [mul_self_star_comm, one_add_self_mul_star_eq_star] - -lemma star_mul_le_right_of_mul_le_right (l m : language α) : l * m ≤ m → l.star * m ≤ m := -begin - intro h, - rw [star_eq_supr_pow, supr_mul], - refine supr_le _, - intro n, - induction n with n ih, - { simp }, - rw [pow_succ', mul_assoc (l^n) l m], - exact le_trans (le_mul_congr le_rfl h) ih, -end - -lemma star_mul_le_left_of_mul_le_left (l m : language α) : m * l ≤ m → m * l.star ≤ m := -begin - intro h, - rw [star_eq_supr_pow, mul_supr], - refine supr_le _, - intro n, - induction n with n ih, - { simp }, - rw [pow_succ, ←mul_assoc m l (l^n)], - exact le_trans (le_mul_congr h le_rfl) ih -end +@[simp] lemma one_add_kstar_mul_self_eq_kstar (l : language α) : 1 + l∗ * l = l∗ := +by rw [mul_self_kstar_comm, one_add_self_mul_kstar_eq_kstar] + +instance : kleene_algebra (language α) := +{ one_le_kstar := λ a l hl, ⟨[], hl, by simp⟩, + mul_kstar_le_kstar := λ a, (one_add_self_mul_kstar_eq_kstar a).le.trans' le_sup_right, + kstar_mul_le_kstar := λ a, (one_add_kstar_mul_self_eq_kstar a).le.trans' le_sup_right, + kstar_mul_le_self := λ l m h, begin + rw [kstar_eq_supr_pow, supr_mul], + refine supr_le (λ n, _), + induction n with n ih, + { simp }, + rw [pow_succ', mul_assoc (l^n) l m], + exact le_trans (le_mul_congr le_rfl h) ih, + end, + mul_kstar_le_self := λ l m h, begin + rw [kstar_eq_supr_pow, mul_supr], + refine supr_le (λ n, _), + induction n with n ih, + { simp }, + rw [pow_succ, ←mul_assoc m l (l^n)], + exact le_trans (le_mul_congr h le_rfl) ih, + end, + ..language.semiring, ..set.complete_boolean_algebra, ..language.has_kstar } end language diff --git a/src/computability/regular_expressions.lean b/src/computability/regular_expressions.lean index bcd047edbcfd6..d1f36e987068e 100644 --- a/src/computability/regular_expressions.lean +++ b/src/computability/regular_expressions.lean @@ -20,6 +20,7 @@ computer science such as the POSIX standard. -/ open list set +open_locale computability universe u @@ -69,7 +70,7 @@ attribute [pattern] has_mul.mul | (char a) := {[a]} | (P + Q) := P.matches + Q.matches | (P * Q) := P.matches * Q.matches -| (star P) := P.matches.star +| (star P) := P.matches∗ @[simp] lemma matches_zero : (0 : regular_expression α).matches = 0 := rfl @[simp] lemma matches_epsilon : (1 : regular_expression α).matches = 1 := rfl @@ -82,7 +83,7 @@ attribute [pattern] has_mul.mul ∀ n : ℕ, (P ^ n).matches = P.matches ^ n | 0 := matches_epsilon | (n + 1) := (matches_mul _ _).trans $ eq.trans (congr_arg _ (matches_pow n)) (pow_succ _ _).symm -@[simp] lemma matches_star (P : regular_expression α) : P.star.matches = P.matches.star := rfl +@[simp] lemma matches_star (P : regular_expression α) : P.star.matches = P.matches∗ := rfl /-- `match_epsilon P` is true if and only if `P` matches the empty string -/ def match_epsilon : regular_expression α → bool @@ -299,7 +300,7 @@ begin rw ←ih₂ at hmatch₂, exact ⟨ x, y, hsum.symm, hmatch₁, hmatch₂ ⟩ } }, case star : _ ih - { rw [star_rmatch_iff, language.star_def_nonempty], + { rw [star_rmatch_iff, language.kstar_def_nonempty], split, all_goals { rintro ⟨ S, hx, hS ⟩, @@ -363,7 +364,7 @@ omit dec | (R * S) := by simp only [matches_map, map, matches_mul, map_mul] | (star R) := begin simp_rw [map, matches, matches_map], - rw [language.star_eq_supr_pow, language.star_eq_supr_pow], + rw [language.kstar_eq_supr_pow, language.kstar_eq_supr_pow], simp_rw ←map_pow, exact image_Union.symm, end From 168eeff2a36ad2298744206421f5d33784e9b3d4 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9mi=20Bottinelli?= Date: Wed, 25 Jan 2023 14:10:25 +0000 Subject: [PATCH 0355/1291] feat(analysis/constant_speed): define constant speed functions and prove existence of the arc-length reparameterization (#18208) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Junyan Xu Co-authored-by: Rémi Bottinelli --- src/analysis/bounded_variation.lean | 155 ++++++++++++--- src/analysis/constant_speed.lean | 290 ++++++++++++++++++++++++++++ 2 files changed, 421 insertions(+), 24 deletions(-) create mode 100644 src/analysis/constant_speed.lean diff --git a/src/analysis/bounded_variation.lean b/src/analysis/bounded_variation.lean index aa78a0f8eda05..5809e1b74a97b 100644 --- a/src/analysis/bounded_variation.lean +++ b/src/analysis/bounded_variation.lean @@ -7,6 +7,8 @@ import measure_theory.measure.lebesgue import analysis.calculus.monotone import data.set.function import algebra.group.basic +import tactic.swap_var +import tactic.wlog /-! # Functions of bounded variation @@ -86,7 +88,7 @@ begin rw [edist_congr_right (h $ p.snd.prop.2 (i+1)), edist_congr_left (h $ p.snd.prop.2 i)], end -lemma eq_of_eq_on {f f' : α → E} {s : set α} (h : set.eq_on f f' s) : +lemma eq_of_eq_on {f f' : α → E} {s : set α} (h : eq_on f f' s) : evariation_on f s = evariation_on f' s := eq_of_edist_zero_on (λ x xs, by rw [h xs, edist_self]) @@ -590,13 +592,13 @@ begin end lemma comp_le_of_monotone_on (f : α → E) {s : set α} {t : set β} (φ : β → α) - (hφ : monotone_on φ t) (φst : set.maps_to φ t s) : + (hφ : monotone_on φ t) (φst : maps_to φ t s) : evariation_on (f ∘ φ) t ≤ evariation_on f s := supr_le $ λ ⟨n, u, hu, ut⟩, le_supr_of_le ⟨n, φ ∘ u, λ x y xy, hφ (ut x) (ut y) (hu xy), λ i, φst (ut i)⟩ le_rfl lemma comp_le_of_antitone_on (f : α → E) {s : set α} {t : set β} (φ : β → α) - (hφ : antitone_on φ t) (φst : set.maps_to φ t s) : + (hφ : antitone_on φ t) (φst : maps_to φ t s) : evariation_on (f ∘ φ) t ≤ evariation_on f s := begin refine supr_le _, @@ -609,39 +611,73 @@ begin simpa only [tsub_pos_iff_lt, finset.mem_range] using hx, end -lemma comp_eq_of_monotone_on (f : α → E) {s : set α} {t : set β} (φ : β → α) - (hφ : monotone_on φ t) (φst : set.maps_to φ t s) (φsur : set.surj_on φ t s) : - evariation_on (f ∘ φ) t = evariation_on f s := +lemma comp_eq_of_monotone_on (f : α → E) {t : set β} (φ : β → α) (hφ : monotone_on φ t) : + evariation_on (f ∘ φ) t = evariation_on f (φ '' t) := begin - apply le_antisymm (comp_le_of_monotone_on f φ hφ φst), + apply le_antisymm (comp_le_of_monotone_on f φ hφ (maps_to_image φ t)), casesI is_empty_or_nonempty β, { convert zero_le _, - exact evariation_on.subsingleton f ((subsingleton_of_subsingleton.image _).anti φsur) }, + exact evariation_on.subsingleton f + ((subsingleton_of_subsingleton.image _).anti (surj_on_image φ t)) }, let ψ := φ.inv_fun_on t, - have ψφs : set.eq_on (φ ∘ ψ) id s := φsur.right_inv_on_inv_fun_on, - have ψts : set.maps_to ψ s t := φsur.maps_to_inv_fun_on, - have hψ : monotone_on ψ s := + have ψφs : eq_on (φ ∘ ψ) id (φ '' t) := (surj_on_image φ t).right_inv_on_inv_fun_on, + have ψts : maps_to ψ (φ '' t) t := (surj_on_image φ t).maps_to_inv_fun_on, + have hψ : monotone_on ψ (φ '' t) := function.monotone_on_of_right_inv_on_of_maps_to hφ ψφs ψts, - change evariation_on (f ∘ id) s ≤ evariation_on (f ∘ φ) t, - rw ←eq_of_eq_on (ψφs.comp_left : set.eq_on (f ∘ (φ ∘ ψ)) (f ∘ id) s), + change evariation_on (f ∘ id) (φ '' t) ≤ evariation_on (f ∘ φ) t, + rw ←eq_of_eq_on (ψφs.comp_left : eq_on (f ∘ (φ ∘ ψ)) (f ∘ id) (φ '' t)), exact comp_le_of_monotone_on _ ψ hψ ψts, end -lemma comp_eq_of_antitone_on (f : α → E) {s : set α} {t : set β} (φ : β → α) - (hφ : antitone_on φ t) (φst : set.maps_to φ t s) (φsur : set.surj_on φ t s) : - evariation_on (f ∘ φ) t = evariation_on f s := +-- porting note: move to file `data.set.intervals.basic` once the port is over, +-- and use it in theorem `polynomial_functions_closure_eq_top` +-- in the file `topology/continuous_function/weierstrass.lean` +lemma _root_.set.subsingleton_Icc_of_ge {α : Type*} [partial_order α] {a b : α} (h : b ≤ a) : + set.subsingleton (Icc a b) := begin - apply le_antisymm (comp_le_of_antitone_on f φ hφ φst), + rintros c ⟨ac,cb⟩ d ⟨ad,db⟩, + cases le_antisymm (cb.trans h) ac, + cases le_antisymm (db.trans h) ad, + refl, +end + +lemma comp_inter_Icc_eq_of_monotone_on (f : α → E) {t : set β} (φ : β → α) + (hφ : monotone_on φ t) {x y : β} (hx : x ∈ t) (hy : y ∈ t) : + evariation_on (f ∘ φ) (t ∩ Icc x y) = evariation_on f ((φ '' t) ∩ Icc (φ x) (φ y)) := +begin + rcases le_total x y with h|h, + { convert comp_eq_of_monotone_on f φ (hφ.mono (set.inter_subset_left t (Icc x y))), + apply le_antisymm, + { rintro _ ⟨⟨u, us, rfl⟩, vφx, vφy⟩, + rcases le_total x u with xu|ux, + { rcases le_total u y with uy|yu, + { exact ⟨u, ⟨us, ⟨xu, uy⟩⟩, rfl⟩, }, + { rw le_antisymm vφy (hφ hy us yu), + exact ⟨y, ⟨hy, ⟨h, le_rfl⟩⟩, rfl⟩, }, }, + { rw ←le_antisymm vφx (hφ us hx ux), + exact ⟨x, ⟨hx, ⟨le_rfl, h⟩⟩, rfl⟩, }, }, + { rintro _ ⟨u, ⟨⟨hu, xu, uy⟩, rfl⟩⟩, + refine ⟨⟨u, hu, rfl⟩, ⟨hφ hx hu xu, hφ hu hy uy⟩⟩, }, }, + { rw [evariation_on.subsingleton, evariation_on.subsingleton], + exacts [(set.subsingleton_Icc_of_ge (hφ hy hx h)).anti (set.inter_subset_right _ _), + (set.subsingleton_Icc_of_ge h).anti (set.inter_subset_right _ _)], }, +end + +lemma comp_eq_of_antitone_on (f : α → E) {t : set β} (φ : β → α) (hφ : antitone_on φ t) : + evariation_on (f ∘ φ) t = evariation_on f (φ '' t) := +begin + apply le_antisymm (comp_le_of_antitone_on f φ hφ (maps_to_image φ t)), casesI is_empty_or_nonempty β, { convert zero_le _, - exact evariation_on.subsingleton f ((subsingleton_of_subsingleton.image _).anti φsur) }, + exact evariation_on.subsingleton f + ((subsingleton_of_subsingleton.image _).anti (surj_on_image φ t)) }, let ψ := φ.inv_fun_on t, - have ψφs : set.eq_on (φ ∘ ψ) id s := φsur.right_inv_on_inv_fun_on, - have ψts : set.maps_to ψ s t := φsur.maps_to_inv_fun_on, - have hψ : antitone_on ψ s := + have ψφs : eq_on (φ ∘ ψ) id (φ '' t) := (surj_on_image φ t).right_inv_on_inv_fun_on, + have ψts := (surj_on_image φ t).maps_to_inv_fun_on, + have hψ : antitone_on ψ (φ '' t) := function.antitone_on_of_right_inv_on_of_maps_to hφ ψφs ψts, - change evariation_on (f ∘ id) s ≤ evariation_on (f ∘ φ) t, - rw ←eq_of_eq_on (ψφs.comp_left : set.eq_on (f ∘ (φ ∘ ψ)) (f ∘ id) s), + change evariation_on (f ∘ id) (φ '' t) ≤ evariation_on (f ∘ φ) t, + rw ←eq_of_eq_on (ψφs.comp_left : eq_on (f ∘ (φ ∘ ψ)) (f ∘ id) (φ '' t)), exact comp_le_of_antitone_on _ ψ hψ ψts, end @@ -649,7 +685,10 @@ open order_dual lemma comp_of_dual (f : α → E) (s : set α) : evariation_on (f ∘ of_dual) (of_dual ⁻¹' s) = evariation_on f s := -comp_eq_of_antitone_on f of_dual (λ _ _ _ _, id) (maps_to_preimage _ _) $ λ x hx, ⟨x, hx, rfl⟩ +begin + convert comp_eq_of_antitone_on f of_dual (λ _ _ _ _, id), + simp only [equiv.image_preimage], +end end evariation_on @@ -753,6 +792,62 @@ begin evariation_on.Icc_add_Icc f xy yz ys], }, end +@[protected] +lemma edist_zero_of_eq_zero + {f : α → E} {s : set α} (hf : has_locally_bounded_variation_on f s) + {a b : α} (ha : a ∈ s) (hb : b ∈ s) (h : variation_on_from_to f s a b = 0) : + edist (f a) (f b) = 0 := +begin + wlog h' : a ≤ b := le_total a b using [b a, a b] tactic.skip, + { apply le_antisymm _ (zero_le _), + rw [←ennreal.of_real_zero, ←h, eq_of_le f s h', ennreal.of_real_to_real (hf a b ha hb)], + apply evariation_on.edist_le, + exacts [⟨ha, ⟨le_rfl, h'⟩⟩, ⟨hb, ⟨h', le_rfl⟩⟩] }, + { assume ha hb hab, + rw edist_comm, + apply this hb ha, + rw [eq_neg_swap, hab, neg_zero] } +end + +@[protected] +lemma eq_left_iff + {f : α → E} {s : set α} (hf : has_locally_bounded_variation_on f s) + {a b c : α} (ha : a ∈ s) (hb : b ∈ s) (hc : c ∈ s) : + variation_on_from_to f s a b = variation_on_from_to f s a c ↔ variation_on_from_to f s b c = 0 := +by simp only [←add hf ha hb hc, self_eq_add_right] + +@[protected] +lemma eq_zero_iff_of_le + {f : α → E} {s : set α} (hf : has_locally_bounded_variation_on f s) + {a b : α} (ha : a ∈ s) (hb : b ∈ s) (ab : a ≤ b) : + variation_on_from_to f s a b = 0 ↔ + ∀ ⦃x⦄ (hx : x ∈ s ∩ Icc a b) ⦃y⦄ (hy : y ∈ s ∩ Icc a b), edist (f x) (f y) = 0 := +by rw [eq_of_le _ _ ab, ennreal.to_real_eq_zero_iff, + or_iff_left (hf a b ha hb), evariation_on.eq_zero_iff] + +@[protected] +lemma eq_zero_iff_of_ge + {f : α → E} {s : set α} (hf : has_locally_bounded_variation_on f s) + {a b : α} (ha : a ∈ s) (hb : b ∈ s) (ba : b ≤ a) : + variation_on_from_to f s a b = 0 ↔ + ∀ ⦃x⦄ (hx : x ∈ s ∩ Icc b a) ⦃y⦄ (hy : y ∈ s ∩ Icc b a), edist (f x) (f y) = 0 := +by rw [eq_of_ge _ _ ba, neg_eq_zero, ennreal.to_real_eq_zero_iff, + or_iff_left (hf b a hb ha), evariation_on.eq_zero_iff] + +@[protected] +lemma eq_zero_iff + {f : α → E} {s : set α} (hf : has_locally_bounded_variation_on f s) + {a b : α} (ha : a ∈ s) (hb : b ∈ s) : + variation_on_from_to f s a b = 0 ↔ + ∀ ⦃x⦄ (hx : x ∈ s ∩ uIcc a b) ⦃y⦄ (hy : y ∈ s ∩ uIcc a b), edist (f x) (f y) = 0 := +begin + rcases le_total a b with ab|ba, + { rw uIcc_of_le ab, + exact eq_zero_iff_of_le hf ha hb ab, }, + { rw uIcc_of_ge ba, + exact eq_zero_iff_of_ge hf ha hb ba, }, +end + variables {f} {s} @[protected] @@ -793,6 +888,18 @@ begin by rw [←add hf as bs cs, add_sub_cancel'] end +@[protected] +lemma comp_eq_of_monotone_on (f : α → E) {t : set β} (φ : β → α) (hφ : monotone_on φ t) + {x y : β} (hx : x ∈ t) (hy : y ∈ t) : + variation_on_from_to (f ∘ φ) t x y = variation_on_from_to f (φ '' t) (φ x) (φ y) := +begin + rcases le_total x y with h|h, + { rw [eq_of_le _ _ h, eq_of_le _ _ (hφ hx hy h), + evariation_on.comp_inter_Icc_eq_of_monotone_on f φ hφ hx hy], }, + { rw [eq_of_ge _ _ h, eq_of_ge _ _ (hφ hy hx h), + evariation_on.comp_inter_Icc_eq_of_monotone_on f φ hφ hy hx], }, +end + end variation_on_from_to /-- If a real valued function has bounded variation on a set, then it is a difference of monotone diff --git a/src/analysis/constant_speed.lean b/src/analysis/constant_speed.lean new file mode 100644 index 0000000000000..ea5e9d545fe72 --- /dev/null +++ b/src/analysis/constant_speed.lean @@ -0,0 +1,290 @@ +/- +Copyright (c) 2023 Rémi Bottinelli. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Rémi Bottinelli +-/ +import measure_theory.measure.lebesgue +import analysis.calculus.monotone +import data.set.function +import analysis.bounded_variation +import tactic.swap_var +/-! +# Constant speed + +This file defines the notion of constant (and unit) speed for a function `f : ℝ → E` with +pseudo-emetric structure on `E` with respect to a set `s : set ℝ` and "speed" `l : ℝ≥0`, and shows +that if `f` has locally bounded variation on `s`, it can be obtained (up to distance zero, on `s`), +as a composite `φ ∘ (variation_on_from_to f s a)`, where `φ` has unit speed and `a ∈ s`. + +## Main definitions + +* `has_constant_speed_on_with f s l`, stating that the speed of `f` on `s` is `l`. +* `has_unit_speed_on f s`, stating that the speed of `f` on `s` is `1`. +* `natural_parameterization f s a : ℝ → E`, the unit speed reparameterization of `f` on `s` relative + to `a`. + +## Main statements + +* `unique_unit_speed_on_Icc_zero` proves that if `f` and `f ∘ φ` are both naturally + parameterized on closed intervals starting at `0`, then `φ` must be the identity on + those intervals. +* `edist_natural_parameterization_eq_zero` proves that if `f` has locally bounded variation, then + precomposing `natural_parameterization f s a` with `variation_on_from_to f s a` yields a function + at distance zero from `f` on `s`. +* `has_unit_speed_natural_parameterization` proves that if `f` has locally bounded + variation, then `natural_parameterization f s a` has unit speed on `s`. + +## Tags + +arc-length, parameterization +-/ + +open_locale big_operators nnreal ennreal +open set measure_theory classical + +variables {α : Type*} [linear_order α] {E : Type*} [pseudo_emetric_space E] +variables (f : ℝ → E) (s : set ℝ) (l : ℝ≥0) + +/-- +`f` has constant speed `l` on `s` if the variation of `f` on `s ∩ Icc x y` is equal to +`l * (y - x)` for any `x y` in `s`. +-/ +def has_constant_speed_on_with := +∀ ⦃x⦄ (hx : x ∈ s) ⦃y⦄ (hy : y ∈ s), evariation_on f (s ∩ Icc x y) = ennreal.of_real (l * (y - x)) + +variables {f} {s} {l} + +lemma has_constant_speed_on_with.has_locally_bounded_variation_on + (h : has_constant_speed_on_with f s l) : has_locally_bounded_variation_on f s := λ x y hx hy, +by simp only [has_bounded_variation_on, h hx hy, ne.def, ennreal.of_real_ne_top, not_false_iff] + +lemma has_constant_speed_on_with_of_subsingleton + (f : ℝ → E) {s : set ℝ} (hs : s.subsingleton) (l : ℝ≥0) : has_constant_speed_on_with f s l := +begin + rintro x hx y hy, cases hs hx hy, + rw evariation_on.subsingleton f (λ y hy z hz, hs hy.1 hz.1 : (s ∩ Icc x x).subsingleton), + simp only [sub_self, mul_zero, ennreal.of_real_zero], +end + +lemma has_constant_speed_on_with_iff_ordered : + has_constant_speed_on_with f s l ↔ + ∀ ⦃x⦄ (hx : x ∈ s) ⦃y⦄ (hy : y ∈ s), (x ≤ y) → + evariation_on f (s ∩ Icc x y) = ennreal.of_real (l * (y - x)) := +begin + refine ⟨λ h x xs y ys xy, h xs ys, λ h x xs y ys, _⟩, + rcases le_total x y with xy|yx, + { exact h xs ys xy, }, + { rw [evariation_on.subsingleton, ennreal.of_real_of_nonpos], + { exact mul_nonpos_of_nonneg_of_nonpos l.prop (sub_nonpos_of_le yx), }, + { rintro z ⟨zs, xz, zy⟩ w ⟨ws, xw, wy⟩, + cases le_antisymm (zy.trans yx) xz, + cases le_antisymm (wy.trans yx) xw, + refl, }, }, +end + +lemma has_constant_speed_on_with_iff_variation_on_from_to_eq : + has_constant_speed_on_with f s l ↔ (has_locally_bounded_variation_on f s ∧ + ∀ ⦃x⦄ (hx : x ∈ s) ⦃y⦄ (hy : y ∈ s), variation_on_from_to f s x y = l * (y - x)) := +begin + split, + { rintro h, refine ⟨h.has_locally_bounded_variation_on, λ x xs y ys, _⟩, + rw has_constant_speed_on_with_iff_ordered at h, + rcases le_total x y with xy|yx, + { rw [variation_on_from_to.eq_of_le f s xy, h xs ys xy, + ennreal.to_real_of_real (mul_nonneg l.prop (sub_nonneg.mpr xy))], }, + { rw [variation_on_from_to.eq_of_ge f s yx, h ys xs yx, + ennreal.to_real_of_real (mul_nonneg l.prop (sub_nonneg.mpr yx)), + mul_comm ↑l, mul_comm ↑l, ←neg_mul, neg_sub], }, }, + { rw has_constant_speed_on_with_iff_ordered, + rintro h x xs y ys xy, + rw [←h.2 xs ys, variation_on_from_to.eq_of_le f s xy, + ennreal.of_real_to_real (h.1 x y xs ys)], }, +end + +lemma has_constant_speed_on_with.union {t : set ℝ} + (hfs : has_constant_speed_on_with f s l) (hft : has_constant_speed_on_with f t l) + {x : ℝ} (hs : is_greatest s x) (ht : is_least t x) : has_constant_speed_on_with f (s ∪ t) l := +begin + rw has_constant_speed_on_with_iff_ordered at hfs hft ⊢, + rintro z (zs|zt) y (ys|yt) zy, + { have : (s ∪ t) ∩ Icc z y = (s ∩ Icc z y), by + { ext w, split, + { rintro ⟨(ws|wt), zw, wy⟩, + { exact ⟨ws, zw, wy⟩, }, + { exact ⟨(le_antisymm (wy.trans (hs.2 ys)) (ht.2 wt)).symm ▸ hs.1, zw, wy⟩, }, }, + { rintro ⟨ws, zwy⟩, exact ⟨or.inl ws, zwy⟩, }, }, + rw [this, hfs zs ys zy], }, + { have : (s ∪ t) ∩ Icc z y = (s ∩ Icc z x) ∪ (t ∩ Icc x y), by + { ext w, split, + { rintro ⟨(ws|wt), zw, wy⟩, + exacts [or.inl ⟨ws, zw, hs.2 ws⟩, or.inr ⟨wt, ht.2 wt, wy⟩], }, + { rintro (⟨ws, zw, wx⟩|⟨wt, xw, wy⟩), + exacts [⟨or.inl ws, zw, wx.trans (ht.2 yt)⟩, ⟨or.inr wt, (hs.2 zs).trans xw, wy⟩], }, }, + rw [this, + @evariation_on.union _ _ _ _ f _ _ x, + hfs zs hs.1 (hs.2 zs), hft ht.1 yt (ht.2 yt), + ←ennreal.of_real_add (mul_nonneg l.prop (sub_nonneg.mpr (hs.2 zs))) + (mul_nonneg l.prop (sub_nonneg.mpr (ht.2 yt))) ], + ring_nf, + exacts [⟨⟨hs.1, hs.2 zs, le_rfl⟩, λ w ⟨ws, zw, wx⟩, wx⟩, + ⟨⟨ht.1, le_rfl, ht.2 yt⟩, λ w ⟨wt, xw, wy⟩, xw⟩], }, + { cases le_antisymm zy ((hs.2 ys).trans (ht.2 zt)), + simp only [Icc_self, sub_self, mul_zero, ennreal.of_real_zero], + exact evariation_on.subsingleton _ (λ _ ⟨_, uz⟩ _ ⟨_, vz⟩, uz.trans vz.symm), }, + { have : (s ∪ t) ∩ Icc z y = (t ∩ Icc z y), by + { ext w, split, + { rintro ⟨(ws|wt), zw, wy⟩, + { exact ⟨(le_antisymm ((ht.2 zt).trans zw) (hs.2 ws)) ▸ ht.1, zw, wy⟩, }, + { exact ⟨wt, zw, wy⟩, }, }, + { rintro ⟨wt, zwy⟩, exact ⟨or.inr wt, zwy⟩, }, }, + rw [this, hft zt yt zy], } +end + +lemma has_constant_speed_on_with.Icc_Icc {x y z : ℝ} + (hfs : has_constant_speed_on_with f (Icc x y) l) + (hft : has_constant_speed_on_with f (Icc y z) l) : has_constant_speed_on_with f (Icc x z) l := +begin + rcases le_total x y with xy|yx, + rcases le_total y z with yz|zy, + { rw ←set.Icc_union_Icc_eq_Icc xy yz, + exact hfs.union hft (is_greatest_Icc xy) (is_least_Icc yz), }, + { rintro u ⟨xu, uz⟩ v ⟨xv, vz⟩, + rw [Icc_inter_Icc, sup_of_le_right xu, inf_of_le_right vz, + ←hfs ⟨xu, uz.trans zy⟩ ⟨xv, vz.trans zy⟩, + Icc_inter_Icc, sup_of_le_right xu, inf_of_le_right (vz.trans zy)], }, + { rintro u ⟨xu, uz⟩ v ⟨xv, vz⟩, + rw [Icc_inter_Icc, sup_of_le_right xu, inf_of_le_right vz, + ←hft ⟨yx.trans xu, uz⟩ ⟨yx.trans xv, vz⟩, + Icc_inter_Icc, sup_of_le_right (yx.trans xu), inf_of_le_right (vz)], }, +end + +lemma has_constant_speed_on_with_zero_iff : + has_constant_speed_on_with f s 0 ↔ ∀ x y ∈ s, edist (f x) (f y) = 0 := +begin + dsimp [has_constant_speed_on_with], + simp only [zero_mul, ennreal.of_real_zero, ←evariation_on.eq_zero_iff], + split, + { by_contra', + obtain ⟨h, hfs⟩ := this, + simp_rw evariation_on.eq_zero_iff at hfs h, + push_neg at hfs, + obtain ⟨x, xs, y, ys, hxy⟩ := hfs, + rcases le_total x y with xy|yx, + { exact hxy (h xs ys x ⟨xs, le_rfl, xy⟩ y ⟨ys, xy, le_rfl⟩), }, + { rw edist_comm at hxy, + exact hxy (h ys xs y ⟨ys, le_rfl, yx⟩ x ⟨xs, yx, le_rfl⟩), }, }, + { rintro h x xs y ys, + refine le_antisymm _ (zero_le'), + rw ←h, + exact evariation_on.mono f (inter_subset_left s (Icc x y)), }, +end + +lemma has_constant_speed_on_with.ratio {l' : ℝ≥0} (hl' : l' ≠ 0) {φ : ℝ → ℝ} + (φm : monotone_on φ s) + (hfφ : has_constant_speed_on_with (f ∘ φ) s l) + (hf : has_constant_speed_on_with f (φ '' s) l') + ⦃x : ℝ⦄ (xs : x ∈ s) : eq_on φ (λ y, (l / l') * (y - x) + (φ x)) s := +begin + rintro y ys, + rw [←sub_eq_iff_eq_add, mul_comm, ←mul_div_assoc, + eq_div_iff (nnreal.coe_ne_zero.mpr hl')], + rw has_constant_speed_on_with_iff_variation_on_from_to_eq at hf, + rw has_constant_speed_on_with_iff_variation_on_from_to_eq at hfφ, + symmetry, + calc (y - x) * l + = l * (y - x) : by rw mul_comm + ... = variation_on_from_to (f ∘ φ) s x y : (hfφ.2 xs ys).symm + ... = variation_on_from_to f (φ '' s) (φ x) (φ y) : + variation_on_from_to.comp_eq_of_monotone_on f φ φm xs ys + ... = l' * (φ y - φ x) : hf.2 ⟨x,xs,rfl⟩ ⟨y,ys,rfl⟩ + ... = (φ y - φ x) * l' : by rw mul_comm, +end + +/-- `f` has unit speed on `s` if it is linearly parameterized by `l = 1` on `s`. -/ +def has_unit_speed_on (f : ℝ → E) (s : set ℝ) := has_constant_speed_on_with f s 1 + +lemma has_unit_speed_on.union {t : set ℝ} {x : ℝ} + (hfs : has_unit_speed_on f s) (hft : has_unit_speed_on f t) + (hs : is_greatest s x) (ht : is_least t x) : has_unit_speed_on f (s ∪ t) := +has_constant_speed_on_with.union hfs hft hs ht + +lemma has_unit_speed_on.Icc_Icc {x y z : ℝ} + (hfs : has_unit_speed_on f (Icc x y)) (hft : has_unit_speed_on f (Icc y z)) : + has_unit_speed_on f (Icc x z) := +has_constant_speed_on_with.Icc_Icc hfs hft + +/-- +If both `f` and `f ∘ φ` have unit speed (on `t` and `s` respectively) and `φ` +monotonically maps `s` onto `t`, then `φ` is just a translation (on `s`). +-/ +lemma unique_unit_speed {φ : ℝ → ℝ} (φm : monotone_on φ s) + (hfφ : has_unit_speed_on (f ∘ φ) s) (hf : has_unit_speed_on f (φ '' s)) + ⦃x : ℝ⦄ (xs : x ∈ s) : eq_on φ (λ y, (y - x) + (φ x)) s := +begin + dsimp only [has_unit_speed_on] at hf hfφ, + convert has_constant_speed_on_with.ratio one_ne_zero φm hfφ hf xs, + simp only [nonneg.coe_one, div_self, ne.def, one_ne_zero, not_false_iff, one_mul], +end + +/-- +If both `f` and `f ∘ φ` have unit speed (on `Icc 0 t` and `Icc 0 s` respectively) +and `φ` monotonically maps `Icc 0 s` onto `Icc 0 t`, then `φ` is the identity on `Icc 0 s` +-/ +lemma unique_unit_speed_on_Icc_zero {s t : ℝ} (hs : 0 ≤ s) (ht : 0 ≤ t) + {φ : ℝ → ℝ} (φm : monotone_on φ $ Icc 0 s) (φst : φ '' (Icc 0 s) = (Icc 0 t)) + (hfφ : has_unit_speed_on (f ∘ φ) (Icc 0 s)) + (hf : has_unit_speed_on f (Icc 0 t)) : eq_on φ id (Icc 0 s) := +begin + rw ←φst at hf, + convert unique_unit_speed φm hfφ hf ⟨le_rfl, hs⟩, + have : φ 0 = 0, by + { obtain ⟨x,xs,hx⟩ := φst.rec_on (surj_on_image φ (Icc 0 s)) ⟨le_rfl, ht⟩, + exact le_antisymm (hx.rec_on (φm ⟨le_rfl,hs⟩ xs xs.1)) + (φst.rec_on (maps_to_image φ (Icc 0 s)) (⟨le_rfl, hs⟩)).1, }, + simp only [tsub_zero, this, add_zero], + refl, +end + +/-- +The natural parameterization of `f` on `s`, which, if `f` has locally bounded variation on `s`, +* has unit speed on `s` + (by `natural_parameterization_has_unit_speed`). +* composed with `variation_on_from_to f s a`, is at distance zero from `f` + (by `natural_parameterization_edist_zero`). +-/ +noncomputable def natural_parameterization (f : α → E) (s : set α) (a : α) : ℝ → E := +f ∘ (@function.inv_fun_on _ _ ⟨a⟩ (variation_on_from_to f s a) s) + +lemma edist_natural_parameterization_eq_zero {f : α → E} {s : set α} + (hf : has_locally_bounded_variation_on f s) {a : α} (as : a ∈ s) {b : α} (bs : b ∈ s) : + edist (natural_parameterization f s a (variation_on_from_to f s a b)) (f b) = 0 := +begin + dsimp only [natural_parameterization], + haveI : nonempty α := ⟨a⟩, + let c := function.inv_fun_on (variation_on_from_to f s a) s (variation_on_from_to f s a b), + obtain ⟨cs, hc⟩ := @function.inv_fun_on_pos _ _ _ s + (variation_on_from_to f s a) (variation_on_from_to f s a b) ⟨b, bs, rfl⟩, + rw [variation_on_from_to.eq_left_iff hf as cs bs] at hc, + apply variation_on_from_to.edist_zero_of_eq_zero hf cs bs hc, +end + +lemma has_unit_speed_natural_parameterization (f : α → E) {s : set α} + (hf : has_locally_bounded_variation_on f s) {a : α} (as : a ∈ s) : + has_unit_speed_on (natural_parameterization f s a) (variation_on_from_to f s a '' s) := +begin + dsimp only [has_unit_speed_on], + rw has_constant_speed_on_with_iff_ordered, + rintro _ ⟨b, bs, rfl⟩ _ ⟨c, cs, rfl⟩ h, + rcases le_total c b with cb|bc, + { rw [nnreal.coe_one, one_mul, le_antisymm h (variation_on_from_to.monotone_on hf as cs bs cb), + sub_self, ennreal.of_real_zero, Icc_self, evariation_on.subsingleton], + exact λ x hx y hy, hx.2.trans hy.2.symm, }, + { rw [nnreal.coe_one, one_mul, sub_eq_add_neg, variation_on_from_to.eq_neg_swap, neg_neg, + add_comm, variation_on_from_to.add hf bs as cs, ←variation_on_from_to.eq_neg_swap f], + rw [←evariation_on.comp_inter_Icc_eq_of_monotone_on (natural_parameterization f s a) _ + (variation_on_from_to.monotone_on hf as) bs cs], + rw [@evariation_on.eq_of_edist_zero_on _ _ _ _ _ f], + { rw [variation_on_from_to.eq_of_le _ _ bc, ennreal.of_real_to_real (hf b c bs cs)], }, + { rintro x ⟨xs, bx, xc⟩, + exact edist_natural_parameterization_eq_zero hf as xs, }, }, +end From 996b0ff959da753a555053a480f36e5f264d4207 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Wed, 25 Jan 2023 14:10:27 +0000 Subject: [PATCH 0356/1291] feat(set_theory/zfc/basic): more basic simp/ext lemmas (#18233) We prove class extensionality, the characterization of class unions, and other very simple results. --- src/set_theory/zfc/basic.lean | 29 +++++++++++++++++++++++++++++ 1 file changed, 29 insertions(+) diff --git a/src/set_theory/zfc/basic.lean b/src/set_theory/zfc/basic.lean index 975ff49883ea7..6482af029f56a 100644 --- a/src/set_theory/zfc/basic.lean +++ b/src/set_theory/zfc/basic.lean @@ -677,6 +677,8 @@ theorem singleton_injective : function.injective (@singleton Set Set _) := @[simp] theorem to_set_sUnion (x : Set.{u}) : (⋃₀ x).to_set = ⋃₀ (to_set '' x.to_set) := by { ext, simp } +@[simp] theorem sUnion_empty : ⋃₀ (∅ : Set.{u}) = ∅ := by { ext, simp } + /-- The binary union operation -/ protected def union (x y : Set.{u}) : Set.{u} := ⋃₀ {x, y} @@ -874,6 +876,8 @@ def to_Set (B : Class.{u}) (A : Class.{u}) : Prop := ∃ x, ↑x = A ∧ B x protected def mem (A B : Class.{u}) : Prop := to_Set.{u} B A instance : has_mem Class Class := ⟨Class.mem⟩ +@[simp] theorem not_mem_empty (x : Class.{u}) : x ∉ (∅ : Class.{u}) := λ ⟨_, _, h⟩, h + theorem mem_univ {A : Class.{u}} : A ∈ univ.{u} ↔ ∃ x : Set.{u}, ↑x = A := exists_congr $ λ x, and_true _ @@ -953,6 +957,31 @@ set.ext $ λ z, iff.symm Set.mem_powerset set.ext $ λ z, by { refine iff.trans _ Set.mem_sUnion.symm, exact ⟨λ ⟨._, ⟨a, rfl, ax⟩, za⟩, ⟨a, ax, za⟩, λ ⟨a, ax, za⟩, ⟨_, ⟨a, rfl, ax⟩, za⟩⟩ } +@[ext] theorem ext {x y : Class.{u}} : (∀ z : Class.{u}, z ∈ x ↔ z ∈ y) → x = y := +begin + refine λ h, set.ext (λ z, _), + change x z ↔ y z, + rw [←mem_hom_left z x, ←mem_hom_left z y], + exact h z +end + +theorem ext_iff {x y : Class.{u}} : x = y ↔ (∀ z : Class.{u}, z ∈ x ↔ z ∈ y) := +⟨λ h, by simp [h], ext⟩ + +theorem coe_mem_powerset {x : Class.{u}} {y : Set.{u}} : powerset x y ↔ ↑y ⊆ x := iff.rfl + +@[simp] theorem mem_sUnion {x y : Class.{u}} : y ∈ ⋃₀ x ↔ ∃ z, z ∈ x ∧ y ∈ z := +begin + split, + { rintro ⟨w, rfl, ⟨z, hzx, hwz⟩⟩, + exact ⟨z, hzx, (mem_hom_left _ _).2 hwz⟩ }, + { rintro ⟨w, hwx, ⟨z, rfl, hwz⟩⟩, + exact ⟨z, rfl, ⟨w, hwx, hwz⟩⟩ } +end + +@[simp] theorem sUnion_empty : ⋃₀ (∅ : Class.{u}) = (∅ : Class.{u}) := +by { ext, simp } + /-- The definite description operator, which is `{x}` if `{y | A y} = {x}` and `∅` otherwise. -/ def iota (A : Class) : Class := ⋃₀ {x | ∀ y, A y ↔ y = x} From 27f315c5591c84687852f816d8ef31fe103d03de Mon Sep 17 00:00:00 2001 From: Floris van Doorn Date: Wed, 25 Jan 2023 16:34:38 +0000 Subject: [PATCH 0357/1291] feat(analysis/calculus/cont_diff): Prove that `fderiv_within` is `C^n` for functions with parameters (#16946) * Prove that `fderiv_within` is `C^n` (at a point within a set) for a function with parameters * There are some inconvenient side-conditions needed for the lemmas, feel free to recommend improvements * `set.diag_image` is not used, but a (useful) left-over used before a refactor. * This is useful for lemmas about `mfderiv` and the smooth vector bundle refactor * From the sphere eversion project --- src/analysis/calculus/cont_diff.lean | 201 +++++++++++++++++++++++---- src/data/set/prod.lean | 9 ++ 2 files changed, 181 insertions(+), 29 deletions(-) diff --git a/src/analysis/calculus/cont_diff.lean b/src/analysis/calculus/cont_diff.lean index 287b83ddb8304..68f3a7110ce41 100644 --- a/src/analysis/calculus/cont_diff.lean +++ b/src/analysis/calculus/cont_diff.lean @@ -169,7 +169,7 @@ universes u v w local attribute [instance, priority 1001] normed_add_comm_group.to_add_comm_group normed_space.to_module' add_comm_group.to_add_comm_monoid -open set fin filter +open set fin filter function open_locale topological_space variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] @@ -177,7 +177,7 @@ variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] {F : Type*} [normed_add_comm_group F] [normed_space 𝕜 F] {G : Type*} [normed_add_comm_group G] [normed_space 𝕜 G] {X : Type*} [normed_add_comm_group X] [normed_space 𝕜 X] -{s s₁ t u : set E} {f f₁ : E → F} {g : F → G} {x : E} {c : F} +{s s₁ t u : set E} {f f₁ : E → F} {g : F → G} {x x₀ : E} {c : F} {b : E × F → G} {m n : ℕ∞} /-! ### Functions with a Taylor series on a domain -/ @@ -2289,6 +2289,22 @@ lemma cont_diff_on.clm_comp {g : X → F →L[𝕜] G} {f : X → E →L[𝕜] F cont_diff_on 𝕜 n (λ x, (g x).comp (f x)) s := is_bounded_bilinear_map_comp.cont_diff.comp_cont_diff_on₂ hg hf +lemma cont_diff.clm_apply {f : E → F →L[𝕜] G} {g : E → F} {n : ℕ∞} + (hf : cont_diff 𝕜 n f) (hg : cont_diff 𝕜 n g) : + cont_diff 𝕜 n (λ x, (f x) (g x)) := +is_bounded_bilinear_map_apply.cont_diff.comp₂ hf hg + +lemma cont_diff_on.clm_apply {f : E → F →L[𝕜] G} {g : E → F} {n : ℕ∞} + (hf : cont_diff_on 𝕜 n f s) (hg : cont_diff_on 𝕜 n g s) : + cont_diff_on 𝕜 n (λ x, (f x) (g x)) s := +is_bounded_bilinear_map_apply.cont_diff.comp_cont_diff_on₂ hf hg + +lemma cont_diff.smul_right {f : E → F →L[𝕜] 𝕜} {g : E → G} {n : ℕ∞} + (hf : cont_diff 𝕜 n f) (hg : cont_diff 𝕜 n g) : + cont_diff 𝕜 n (λ x, (f x).smul_right (g x)) := +-- giving the following implicit type arguments speeds up elaboration significantly +(@is_bounded_bilinear_map_smul_right 𝕜 _ F _ _ G _ _).cont_diff.comp₂ hf hg + end specific_bilinear_maps /-- @@ -2309,44 +2325,171 @@ Warning: see remarks attached to `cont_diff_prod_assoc` lemma cont_diff_prod_assoc_symm : cont_diff 𝕜 ⊤ $ (equiv.prod_assoc E F G).symm := (linear_isometry_equiv.prod_assoc 𝕜 E F G).symm.cont_diff -/-! ### Bundled derivatives -/ - -lemma cont_diff_within_at.fderiv_within' - (hf : cont_diff_within_at 𝕜 n f s x) (hs : ∀ᶠ y in 𝓝[insert x s] x, unique_diff_within_at 𝕜 s y) - (hmn : m + 1 ≤ n) : - cont_diff_within_at 𝕜 m (fderiv_within 𝕜 f s) s x := -begin - have : ∀ k : ℕ, (k + 1 : ℕ∞) ≤ n → cont_diff_within_at 𝕜 k (fderiv_within 𝕜 f s) s x, - { intros k hkn, +/-! ### Bundled derivatives are smooth -/ + +/-- One direction of `cont_diff_within_at_succ_iff_has_fderiv_within_at`, but where all derivatives + are taken within the same set. Version for partial derivatives / functions with parameters. + If `f x` is a `C^n+1` family of functions and `g x` is a `C^n` family of points, then the + derivative of `f x` at `g x` depends in a `C^n` way on `x`. We give a general version of this fact + relative to sets which may not have unique derivatives, in the following form. + If `f : E × F → G` is `C^n+1` at `(x₀, g(x₀))` in `(s ∪ {x₀}) × t ⊆ E × F` and `g : E → F` is + `C^n` at `x₀` within some set `s ⊆ E`, then there is a function `f' : E → F →L[𝕜] G` + that is `C^n` at `x₀` within `s` such that for all `x` sufficiently close to `x₀` within + `s ∪ {x₀}` the function `y ↦ f x y` has derivative `f' x` at `g x` within `t ⊆ F`. + For convenience, we return an explicit set of `x`'s where this holds that is a subset of + `s ∪ {x₀}`. + We need one additional condition, namely that `t` is a neighborhood of `g(x₀)` within `g '' s`. + -/ +lemma cont_diff_within_at.has_fderiv_within_at_nhds {f : E → F → G} {g : E → F} + {t : set F} {n : ℕ} {x₀ : E} + (hf : cont_diff_within_at 𝕜 (n+1) (uncurry f) (insert x₀ s ×ˢ t) (x₀, g x₀)) + (hg : cont_diff_within_at 𝕜 n g s x₀) + (hgt : t ∈ 𝓝[g '' s] g x₀) : + ∃ v ∈ 𝓝[insert x₀ s] x₀, v ⊆ insert x₀ s ∧ ∃ f' : E → F →L[𝕜] G, + (∀ x ∈ v, has_fderiv_within_at (f x) (f' x) t (g x)) ∧ + cont_diff_within_at 𝕜 n (λ x, f' x) s x₀ := +begin + have hst : insert x₀ s ×ˢ t ∈ 𝓝[(λ x, (x, g x)) '' s] (x₀, g x₀), + { refine nhds_within_mono _ _ (nhds_within_prod self_mem_nhds_within hgt), + simp_rw [image_subset_iff, mk_preimage_prod, preimage_id', subset_inter_iff, subset_insert, + true_and, subset_preimage_image] }, + obtain ⟨v, hv, hvs, f', hvf', hf'⟩ := cont_diff_within_at_succ_iff_has_fderiv_within_at'.mp hf, + refine ⟨(λ z, (z, g z)) ⁻¹' v ∩ insert x₀ s, _, inter_subset_right _ _, + λ z, (f' (z, g z)).comp (continuous_linear_map.inr 𝕜 E F), _, _⟩, + { refine inter_mem _ self_mem_nhds_within, + have := mem_of_mem_nhds_within (mem_insert _ _) hv, + refine mem_nhds_within_insert.mpr ⟨this, _⟩, + refine (continuous_within_at_id.prod hg.continuous_within_at).preimage_mem_nhds_within' _, + rw [← nhds_within_le_iff] at hst hv ⊢, + refine (hst.trans $ nhds_within_mono _ $ subset_insert _ _).trans hv }, + { intros z hz, + have := hvf' (z, g z) hz.1, + refine this.comp _ (has_fderiv_at_prod_mk_right _ _).has_fderiv_within_at _, + exact maps_to'.mpr (image_prod_mk_subset_prod_right hz.2) }, + { exact (hf'.continuous_linear_map_comp $ (continuous_linear_map.compL 𝕜 F (E × F) G).flip + (continuous_linear_map.inr 𝕜 E F)).comp_of_mem x₀ + (cont_diff_within_at_id.prod hg) hst }, +end + +/-- The most general lemma stating that `x ↦ fderiv_within 𝕜 (f x) t (g x)` is `C^n` +at a point within a set. +To show that `x ↦ D_yf(x,y)g(x)` (taken within `t`) is `C^m` at `x₀` within `s`, we require that +* `f` is `C^n` at `(x₀, g(x₀))` within `(s ∪ {x₀}) × t` for `n ≥ m+1`. +* `g` is `C^m` at `x₀` within `s`; +* Derivatives are unique at `g(x)` within `t` for `x` sufficiently close to `x₀` within `s ∪ {x₀}`; +* `t` is a neighborhood of `g(x₀)` within `g '' s`; -/ +lemma cont_diff_within_at.fderiv_within'' {f : E → F → G} {g : E → F} + {t : set F} {n : ℕ∞} + (hf : cont_diff_within_at 𝕜 n (function.uncurry f) (insert x₀ s ×ˢ t) (x₀, g x₀)) + (hg : cont_diff_within_at 𝕜 m g s x₀) + (ht : ∀ᶠ x in 𝓝[insert x₀ s] x₀, unique_diff_within_at 𝕜 t (g x)) + (hmn : m + 1 ≤ n) + (hgt : t ∈ 𝓝[g '' s] g x₀) : + cont_diff_within_at 𝕜 m (λ x, fderiv_within 𝕜 (f x) t (g x)) s x₀ := +begin + have : ∀ k : ℕ, (k : ℕ∞) ≤ m → + cont_diff_within_at 𝕜 k (λ x, fderiv_within 𝕜 (f x) t (g x)) s x₀, + { intros k hkm, obtain ⟨v, hv, -, f', hvf', hf'⟩ := - cont_diff_within_at_succ_iff_has_fderiv_within_at'.mp (hf.of_le hkn), - apply hf'.congr_of_eventually_eq_insert, - filter_upwards [hv, hs], + (hf.of_le $ (add_le_add_right hkm 1).trans hmn).has_fderiv_within_at_nhds (hg.of_le hkm) hgt, + refine hf'.congr_of_eventually_eq_insert _, + filter_upwards [hv, ht], exact λ y hy h2y, (hvf' y hy).fderiv_within h2y }, induction m using with_top.rec_top_coe, { obtain rfl := eq_top_iff.mpr hmn, rw [cont_diff_within_at_top], exact λ m, this m le_top }, - exact this m hmn + exact this m le_rfl +end + +/-- A special case of `cont_diff_within_at.fderiv_within''` where we require that `s ⊆ g⁻¹(t)`. -/ +lemma cont_diff_within_at.fderiv_within' {f : E → F → G} {g : E → F} + {t : set F} {n : ℕ∞} + (hf : cont_diff_within_at 𝕜 n (function.uncurry f) (insert x₀ s ×ˢ t) (x₀, g x₀)) + (hg : cont_diff_within_at 𝕜 m g s x₀) + (ht : ∀ᶠ x in 𝓝[insert x₀ s] x₀, unique_diff_within_at 𝕜 t (g x)) + (hmn : m + 1 ≤ n) + (hst : s ⊆ g ⁻¹' t) : + cont_diff_within_at 𝕜 m (λ x, fderiv_within 𝕜 (f x) t (g x)) s x₀ := +hf.fderiv_within'' hg ht hmn $ mem_of_superset self_mem_nhds_within $ image_subset_iff.mpr hst + +/-- A special case of `cont_diff_within_at.fderiv_within'` where we require that `x₀ ∈ s` and there + are unique derivatives everywhere within `t`. -/ +lemma cont_diff_within_at.fderiv_within {f : E → F → G} {g : E → F} + {t : set F} {n : ℕ∞} + (hf : cont_diff_within_at 𝕜 n (function.uncurry f) (s ×ˢ t) (x₀, g x₀)) + (hg : cont_diff_within_at 𝕜 m g s x₀) + (ht : unique_diff_on 𝕜 t) + (hmn : m + 1 ≤ n) (hx₀ : x₀ ∈ s) + (hst : s ⊆ g ⁻¹' t) : + cont_diff_within_at 𝕜 m (λ x, fderiv_within 𝕜 (f x) t (g x)) s x₀ := +begin + rw [← insert_eq_self.mpr hx₀] at hf, + refine hf.fderiv_within' hg _ hmn hst, + rw [insert_eq_self.mpr hx₀], + exact eventually_of_mem self_mem_nhds_within (λ x hx, ht _ (hst hx)) +end + +/-- `x ↦ fderiv_within 𝕜 (f x) t (g x) (k x)` is smooth at a point within a set. -/ +lemma cont_diff_within_at.fderiv_within_apply {f : E → F → G} {g k : E → F} + {t : set F} {n : ℕ∞} + (hf : cont_diff_within_at 𝕜 n (function.uncurry f) (s ×ˢ t) (x₀, g x₀)) + (hg : cont_diff_within_at 𝕜 m g s x₀) + (hk : cont_diff_within_at 𝕜 m k s x₀) + (ht : unique_diff_on 𝕜 t) + (hmn : m + 1 ≤ n) (hx₀ : x₀ ∈ s) + (hst : s ⊆ g ⁻¹' t) : + cont_diff_within_at 𝕜 m (λ x, fderiv_within 𝕜 (f x) t (g x) (k x)) s x₀ := +(cont_diff_fst.clm_apply cont_diff_snd).cont_diff_at.comp_cont_diff_within_at x₀ + ((hf.fderiv_within hg ht hmn hx₀ hst).prod hk) + +/-- `fderiv_within 𝕜 f s` is smooth at `x₀` within `s`. -/ +lemma cont_diff_within_at.fderiv_within_right + (hf : cont_diff_within_at 𝕜 n f s x₀) (hs : unique_diff_on 𝕜 s) + (hmn : (m + 1 : ℕ∞) ≤ n) (hx₀s : x₀ ∈ s) : + cont_diff_within_at 𝕜 m (fderiv_within 𝕜 f s) s x₀ := +cont_diff_within_at.fderiv_within + (cont_diff_within_at.comp (x₀, x₀) hf cont_diff_within_at_snd $ prod_subset_preimage_snd s s) + cont_diff_within_at_id hs hmn hx₀s (by rw [preimage_id']) + +/-- `x ↦ fderiv 𝕜 (f x) (g x)` is smooth at `x₀`. -/ +lemma cont_diff_at.cont_diff_at_fderiv {f : E → F → G} {g : E → F} {n : ℕ∞} + (hf : cont_diff_at 𝕜 n (function.uncurry f) (x₀, g x₀)) + (hg : cont_diff_at 𝕜 m g x₀) + (hmn : m + 1 ≤ n) : + cont_diff_at 𝕜 m (λ x, fderiv 𝕜 (f x) (g x)) x₀ := +begin + simp_rw [← fderiv_within_univ], + refine (cont_diff_within_at.fderiv_within hf.cont_diff_within_at hg.cont_diff_within_at + unique_diff_on_univ hmn (mem_univ x₀) _).cont_diff_at univ_mem, + rw [preimage_univ] end -lemma cont_diff_within_at.fderiv_within - (hf : cont_diff_within_at 𝕜 n f s x) (hs : unique_diff_on 𝕜 s) - (hmn : (m + 1 : ℕ∞) ≤ n) (hxs : x ∈ s) : - cont_diff_within_at 𝕜 m (fderiv_within 𝕜 f s) s x := -hf.fderiv_within' (by { rw [insert_eq_of_mem hxs], exact eventually_of_mem self_mem_nhds_within hs}) - hmn +/-- `x ↦ fderiv 𝕜 (f x) (g x)` is smooth. -/ +lemma cont_diff.fderiv {f : E → F → G} {g : E → F} {n m : ℕ∞} + (hf : cont_diff 𝕜 m $ function.uncurry f) (hg : cont_diff 𝕜 n g) (hnm : n + 1 ≤ m) : + cont_diff 𝕜 n (λ x, fderiv 𝕜 (f x) (g x)) := +cont_diff_iff_cont_diff_at.mpr $ λ x, hf.cont_diff_at.cont_diff_at_fderiv hg.cont_diff_at hnm + +/-- `x ↦ fderiv 𝕜 (f x) (g x)` is continuous. -/ +lemma continuous.fderiv {f : E → F → G} {g : E → F} {n : ℕ∞} + (hf : cont_diff 𝕜 n $ function.uncurry f) (hg : continuous g) (hn : 1 ≤ n) : + continuous (λ x, fderiv 𝕜 (f x) (g x)) := +(hf.fderiv (cont_diff_zero.mpr hg) hn).continuous + +/-- `x ↦ fderiv 𝕜 (f x) (g x) (k x)` is smooth. -/ +lemma cont_diff.fderiv_apply {f : E → F → G} {g k : E → F} {n m : ℕ∞} + (hf : cont_diff 𝕜 m $ function.uncurry f) (hg : cont_diff 𝕜 n g) (hk : cont_diff 𝕜 n k) + (hnm : n + 1 ≤ m) : + cont_diff 𝕜 n (λ x, fderiv 𝕜 (f x) (g x) (k x)) := +(hf.fderiv hg hnm).clm_apply hk /-- The bundled derivative of a `C^{n+1}` function is `C^n`. -/ -lemma cont_diff_on_fderiv_within_apply {m n : with_top ℕ} {s : set E} +lemma cont_diff_on_fderiv_within_apply {m n : ℕ∞} {s : set E} {f : E → F} (hf : cont_diff_on 𝕜 n f s) (hs : unique_diff_on 𝕜 s) (hmn : m + 1 ≤ n) : cont_diff_on 𝕜 m (λp : E × E, (fderiv_within 𝕜 f s p.1 : E →L[𝕜] F) p.2) (s ×ˢ univ) := -have I : cont_diff_on 𝕜 m (λ (x : E), fderiv_within 𝕜 f s x) s := hf.fderiv_within hs hmn, -have J : cont_diff_on 𝕜 m (λ (x : E × E), x.1) (s ×ˢ univ) := cont_diff_fst.cont_diff_on, -have A : cont_diff 𝕜 m (λp : (E →L[𝕜] F) × E, p.1 p.2) := is_bounded_bilinear_map_apply.cont_diff, -have B : cont_diff_on 𝕜 m (λ (p : E × E), ((fderiv_within 𝕜 f s p.fst), p.snd)) (s ×ˢ univ) := -(I.comp J (prod_subset_preimage_fst _ _)).prod is_bounded_linear_map.snd.cont_diff.cont_diff_on, -A.comp_cont_diff_on B +((hf.fderiv_within hs hmn).comp cont_diff_on_fst (prod_subset_preimage_fst _ _)).clm_apply + cont_diff_on_snd /-- If a function is at least `C^1`, its bundled derivative (mapping `(x, v)` to `Df(x) v`) is continuous. -/ @@ -3141,7 +3284,7 @@ lemma cont_diff_on_clm_apply {n : ℕ∞} {f : E → F →L[𝕜] G} {s : set E} [finite_dimensional 𝕜 F] : cont_diff_on 𝕜 n f s ↔ ∀ y, cont_diff_on 𝕜 n (λ x, f x y) s := begin - refine ⟨λ h y, (continuous_linear_map.apply 𝕜 G y).cont_diff.comp_cont_diff_on h, λ h, _⟩, + refine ⟨λ h y, h.clm_apply cont_diff_on_const, λ h, _⟩, let d := finrank 𝕜 F, have hd : d = finrank 𝕜 (fin d → 𝕜) := (finrank_fin_fun 𝕜).symm, let e₁ := continuous_linear_equiv.of_finrank_eq hd, diff --git a/src/data/set/prod.lean b/src/data/set/prod.lean index 03c8cff1d27db..dc29231b9b051 100644 --- a/src/data/set/prod.lean +++ b/src/data/set/prod.lean @@ -353,6 +353,15 @@ prod_subset_iff.trans disjoint_iff_forall_ne.symm lemma diag_preimage_prod_self (s : set α) : (λ x, (x, x)) ⁻¹' (s ×ˢ s) = s := inter_self s +lemma diag_image (s : set α) : (λ x, (x, x)) '' s = diagonal α ∩ (s ×ˢ s) := +begin + ext x, split, + { rintro ⟨x, hx, rfl⟩, exact ⟨rfl, hx, hx⟩ }, + { obtain ⟨x, y⟩ := x, + rintro ⟨rfl : x = y, h2x⟩, + exact mem_image_of_mem _ h2x.1 } +end + end diagonal section off_diag From f93c11933efbc3c2f0299e47b8ff83e9b539cbf6 Mon Sep 17 00:00:00 2001 From: Floris van Doorn Date: Wed, 25 Jan 2023 19:38:11 +0000 Subject: [PATCH 0358/1291] feat(geometry/manifold/mfderiv): more arithmetic (#17793) * generalize most arithmetic from normed algebra to normed space * add some arithmetic with `mfderiv` * redefine `mfderiv` (and variants) to use `if` instead of dependent `if`. --- src/geometry/manifold/cont_mdiff_mfderiv.lean | 4 +- src/geometry/manifold/mfderiv.lean | 131 ++++++++++++++---- 2 files changed, 103 insertions(+), 32 deletions(-) diff --git a/src/geometry/manifold/cont_mdiff_mfderiv.lean b/src/geometry/manifold/cont_mdiff_mfderiv.lean index 22ca9dc12747b..97174c760d71b 100644 --- a/src/geometry/manifold/cont_mdiff_mfderiv.lean +++ b/src/geometry/manifold/cont_mdiff_mfderiv.lean @@ -206,7 +206,7 @@ begin { apply cont_diff_on.prod B _, apply C.congr (λp hp, _), simp only with mfld_simps at hp, - simp only [mfderiv_within, hf.mdifferentiable_on one_le_n _ hp.2, hp.1, dif_pos] + simp only [mfderiv_within, hf.mdifferentiable_on one_le_n _ hp.2, hp.1, if_pos] with mfld_simps }, have D : cont_diff_on 𝕜 m (λ x, (fderiv_within 𝕜 (I' ∘ f ∘ I.symm) (I.symm ⁻¹' s ∩ range I) x)) @@ -633,7 +633,7 @@ begin { exact model_with_corners.unique_diff_at_image I }, { exact differentiable_at_id'.prod (differentiable_at_const _) } }, simp only [tangent_bundle.zero_section, tangent_map, mfderiv, - A, dif_pos, chart_at, basic_smooth_vector_bundle_core.chart, + A, if_pos, chart_at, basic_smooth_vector_bundle_core.chart, basic_smooth_vector_bundle_core.to_vector_bundle_core, tangent_bundle_core, function.comp, continuous_linear_map.map_zero] with mfld_simps, rw ← fderiv_within_inter N (I.unique_diff (I ((chart_at H x) x)) (set.mem_range_self _)) at B, diff --git a/src/geometry/manifold/mfderiv.lean b/src/geometry/manifold/mfderiv.lean index 526119e511de0..ee0b19a28d9ad 100644 --- a/src/geometry/manifold/mfderiv.lean +++ b/src/geometry/manifold/mfderiv.lean @@ -280,7 +280,7 @@ derivative of `f` at `x` within `s`, as a continuous linear map from the tangent tangent space at `f x`. -/ def mfderiv_within (f : M → M') (s : set M) (x : M) : tangent_space I x →L[𝕜] tangent_space I' (f x) := -if h : mdifferentiable_within_at I I' f s x then +if mdifferentiable_within_at I I' f s x then (fderiv_within 𝕜 (written_in_ext_chart_at I I' x f) ((ext_chart_at I x).symm ⁻¹' s ∩ range I) ((ext_chart_at I x) x) : _) else 0 @@ -289,7 +289,7 @@ else 0 `f` at `x`, as a continuous linear map from the tangent space at `x` to the tangent space at `f x`. -/ def mfderiv (f : M → M') (x : M) : tangent_space I x →L[𝕜] tangent_space I' (f x) := -if h : mdifferentiable_at I I' f x then +if mdifferentiable_at I I' f x then (fderiv_within 𝕜 (written_in_ext_chart_at I I' x f : E → E') (range I) ((ext_chart_at I x) x) : _) else 0 @@ -431,11 +431,11 @@ lemma mdifferentiable_within_at_iff_of_mem_source lemma mfderiv_within_zero_of_not_mdifferentiable_within_at (h : ¬ mdifferentiable_within_at I I' f s x) : mfderiv_within I I' f s x = 0 := -by simp only [mfderiv_within, h, dif_neg, not_false_iff] +by simp only [mfderiv_within, h, if_neg, not_false_iff] lemma mfderiv_zero_of_not_mdifferentiable_at (h : ¬ mdifferentiable_at I I' f x) : mfderiv I I' f x = 0 := -by simp only [mfderiv, h, dif_neg, not_false_iff] +by simp only [mfderiv, h, if_neg, not_false_iff] theorem has_mfderiv_within_at.mono (h : has_mfderiv_within_at I I' f t x f') (hst : s ⊆ t) : has_mfderiv_within_at I I' f s x f' := @@ -504,7 +504,7 @@ lemma mdifferentiable_within_at.has_mfderiv_within_at (h : mdifferentiable_withi has_mfderiv_within_at I I' f s x (mfderiv_within I I' f s x) := begin refine ⟨h.1, _⟩, - simp only [mfderiv_within, h, dif_pos] with mfld_simps, + simp only [mfderiv_within, h, if_pos] with mfld_simps, exact differentiable_within_at.has_fderiv_within_at h.2 end @@ -512,20 +512,20 @@ lemma mdifferentiable_within_at.mfderiv_within (h : mdifferentiable_within_at I (mfderiv_within I I' f s x) = fderiv_within 𝕜 (written_in_ext_chart_at I I' x f : _) ((ext_chart_at I x).symm ⁻¹' s ∩ range I) ((ext_chart_at I x) x) := -by simp only [mfderiv_within, h, dif_pos] +by simp only [mfderiv_within, h, if_pos] lemma mdifferentiable_at.has_mfderiv_at (h : mdifferentiable_at I I' f x) : has_mfderiv_at I I' f x (mfderiv I I' f x) := begin refine ⟨h.1, _⟩, - simp only [mfderiv, h, dif_pos] with mfld_simps, + simp only [mfderiv, h, if_pos] with mfld_simps, exact differentiable_within_at.has_fderiv_within_at h.2 end lemma mdifferentiable_at.mfderiv (h : mdifferentiable_at I I' f x) : (mfderiv I I' f x) = fderiv_within 𝕜 (written_in_ext_chart_at I I' x f : _) (range I) ((ext_chart_at I x) x) := -by simp only [mfderiv, h, dif_pos] +by simp only [mfderiv, h, if_pos] lemma has_mfderiv_at.mfderiv (h : has_mfderiv_at I I' f x f') : mfderiv I I' f x = f' := @@ -778,7 +778,7 @@ begin by_cases h : mdifferentiable_within_at I I' f s x, { exact ((h.has_mfderiv_within_at).congr_of_eventually_eq hL hx).mfderiv_within hs }, { unfold mfderiv_within, - rw [dif_neg h, dif_neg], + rw [if_neg h, if_neg], rwa ← hL.mdifferentiable_within_at_iff I I' hx } end @@ -1037,8 +1037,8 @@ alias mdifferentiable_iff_differentiable ↔ mfderiv_within (𝓘(𝕜, E)) (𝓘(𝕜, E')) f s x = fderiv_within 𝕜 f s x := begin by_cases h : mdifferentiable_within_at (𝓘(𝕜, E)) (𝓘(𝕜, E')) f s x, - { simp only [mfderiv_within, h, dif_pos] with mfld_simps }, - { simp only [mfderiv_within, h, dif_neg, not_false_iff], + { simp only [mfderiv_within, h, if_pos] with mfld_simps }, + { simp only [mfderiv_within, h, if_neg, not_false_iff], rw [mdifferentiable_within_at_iff_differentiable_within_at] at h, exact (fderiv_within_zero_of_not_differentiable_within_at h).symm } end @@ -1236,10 +1236,8 @@ Note that in the in `has_mfderiv_at` lemmas there is an abuse of the defeq betwe canonical, but in this case (the tangent space of a vector space) it is canonical. -/ -variables { z : M} {F' : Type*} [normed_comm_ring F'] [normed_algebra 𝕜 F'] -{f g : M → E'} {p q : M → F'} {I} -{f' g' : tangent_space I z →L[𝕜] E'} -{p' q' : tangent_space I z →L[𝕜] F'} +section group +variables {I} {z : M} {f g : M → E'} {f' g' : tangent_space I z →L[𝕜] E'} lemma has_mfderiv_at.add (hf : has_mfderiv_at I 𝓘(𝕜, E') f z f') (hg : has_mfderiv_at I 𝓘(𝕜, E') g z g') : has_mfderiv_at I 𝓘(𝕜, E') (f + g) z (f' + g') := @@ -1253,18 +1251,11 @@ lemma mdifferentiable.add (hf : mdifferentiable I 𝓘(𝕜, E') f) (hg : mdiffe mdifferentiable I 𝓘(𝕜, E') (f + g) := λ x, (hf x).add (hg x) -lemma has_mfderiv_at.mul (hp : has_mfderiv_at I 𝓘(𝕜, F') p z p') - (hq : has_mfderiv_at I 𝓘(𝕜, F') q z q') : - has_mfderiv_at I 𝓘(𝕜, F') (p * q) z (p z • q' + q z • p' : E →L[𝕜] F') := -⟨hp.1.mul hq.1, by simpa only with mfld_simps using hp.2.mul hq.2⟩ - -lemma mdifferentiable_at.mul (hp : mdifferentiable_at I 𝓘(𝕜, F') p z) - (hq : mdifferentiable_at I 𝓘(𝕜, F') q z) : mdifferentiable_at I 𝓘(𝕜, F') (p * q) z := -(hp.has_mfderiv_at.mul hq.has_mfderiv_at).mdifferentiable_at - -lemma mdifferentiable.mul {f g : M → F'} (hf : mdifferentiable I 𝓘(𝕜, F') f) - (hg : mdifferentiable I 𝓘(𝕜, F') g) : mdifferentiable I 𝓘(𝕜, F') (f * g) := -λ x, (hf x).mul (hg x) +lemma mfderiv_add (hf : mdifferentiable_at I 𝓘(𝕜, E') f z) + (hg : mdifferentiable_at I 𝓘(𝕜, E') g z) : + (mfderiv I 𝓘(𝕜, E') (f + g) z : tangent_space I z →L[𝕜] E') = + (mfderiv I 𝓘(𝕜, E') f z + mfderiv I 𝓘(𝕜, E') g z : tangent_space I z →L[𝕜] E') := +(hf.has_mfderiv_at.add hg.has_mfderiv_at).mfderiv lemma has_mfderiv_at.const_smul (hf : has_mfderiv_at I 𝓘(𝕜, E') f z f') (s : 𝕜) : has_mfderiv_at I 𝓘(𝕜, E') (s • f) z (s • f') := @@ -1274,22 +1265,45 @@ lemma mdifferentiable_at.const_smul (hf : mdifferentiable_at I 𝓘(𝕜, E') f mdifferentiable_at I 𝓘(𝕜, E') (s • f) z := (hf.has_mfderiv_at.const_smul s).mdifferentiable_at -lemma mdifferentiable.const_smul {f : M → E'} (s : 𝕜) (hf : mdifferentiable I 𝓘(𝕜, E') f) : +lemma mdifferentiable.const_smul (s : 𝕜) (hf : mdifferentiable I 𝓘(𝕜, E') f) : mdifferentiable I 𝓘(𝕜, E') (s • f) := λ x, (hf x).const_smul s +lemma const_smul_mfderiv (hf : mdifferentiable_at I 𝓘(𝕜, E') f z) (s : 𝕜) : + (mfderiv I 𝓘(𝕜, E') (s • f) z : tangent_space I z →L[𝕜] E') = + (s • mfderiv I 𝓘(𝕜, E') f z : tangent_space I z →L[𝕜] E') := +(hf.has_mfderiv_at.const_smul s).mfderiv + lemma has_mfderiv_at.neg (hf : has_mfderiv_at I 𝓘(𝕜, E') f z f') : has_mfderiv_at I 𝓘(𝕜, E') (-f) z (-f') := ⟨hf.1.neg, hf.2.neg⟩ +lemma has_mfderiv_at_neg : + has_mfderiv_at I 𝓘(𝕜, E') (-f) z (-f') ↔ has_mfderiv_at I 𝓘(𝕜, E') f z f' := +⟨λ hf, by { convert hf.neg; rw [neg_neg] }, λ hf, hf.neg⟩ + lemma mdifferentiable_at.neg (hf : mdifferentiable_at I 𝓘(𝕜, E') f z) : mdifferentiable_at I 𝓘(𝕜, E') (-f) z := hf.has_mfderiv_at.neg.mdifferentiable_at -lemma mdifferentiable.neg {f : M → E'} (hf : mdifferentiable I 𝓘(𝕜, E') f) : +lemma mdifferentiable_at_neg : + mdifferentiable_at I 𝓘(𝕜, E') (-f) z ↔ mdifferentiable_at I 𝓘(𝕜, E') f z := +⟨λ hf, by { convert hf.neg; rw [neg_neg] }, λ hf, hf.neg⟩ + +lemma mdifferentiable.neg (hf : mdifferentiable I 𝓘(𝕜, E') f) : mdifferentiable I 𝓘(𝕜, E') (-f) := λ x, (hf x).neg +lemma mfderiv_neg (f : M → E') (x : M) : + (mfderiv I 𝓘(𝕜, E') (-f) x : tangent_space I x →L[𝕜] E') = + (- mfderiv I 𝓘(𝕜, E') f x : tangent_space I x →L[𝕜] E') := +begin + simp_rw [mfderiv], + by_cases hf : mdifferentiable_at I 𝓘(𝕜, E') f x, + { exact hf.has_mfderiv_at.neg.mfderiv }, + { rw [if_neg hf], rw [← mdifferentiable_at_neg] at hf, rw [if_neg hf, neg_zero] }, +end + lemma has_mfderiv_at.sub (hf : has_mfderiv_at I 𝓘(𝕜, E') f z f') (hg : has_mfderiv_at I 𝓘(𝕜, E') g z g') : has_mfderiv_at I 𝓘(𝕜, E') (f - g) z (f'- g') := ⟨hf.1.sub hg.1, hf.2.sub hg.2⟩ @@ -1298,10 +1312,67 @@ lemma mdifferentiable_at.sub (hf : mdifferentiable_at I 𝓘(𝕜, E') f z) (hg : mdifferentiable_at I 𝓘(𝕜, E') g z) : mdifferentiable_at I 𝓘(𝕜, E') (f - g) z := (hf.has_mfderiv_at.sub hg.has_mfderiv_at).mdifferentiable_at -lemma mdifferentiable.sub {f : M → E'} (hf : mdifferentiable I 𝓘(𝕜, E') f) +lemma mdifferentiable.sub (hf : mdifferentiable I 𝓘(𝕜, E') f) (hg : mdifferentiable I 𝓘(𝕜, E') g) : mdifferentiable I 𝓘(𝕜, E') (f - g) := λ x, (hf x).sub (hg x) +lemma mfderiv_sub (hf : mdifferentiable_at I 𝓘(𝕜, E') f z) + (hg : mdifferentiable_at I 𝓘(𝕜, E') g z) : + (mfderiv I 𝓘(𝕜, E') (f - g) z : tangent_space I z →L[𝕜] E') = + (mfderiv I 𝓘(𝕜, E') f z - mfderiv I 𝓘(𝕜, E') g z : tangent_space I z →L[𝕜] E') := +(hf.has_mfderiv_at.sub hg.has_mfderiv_at).mfderiv + +end group + +section algebra_over_ring +variables {I} {z : M} {F' : Type*} [normed_ring F'] [normed_algebra 𝕜 F'] + {p q : M → F'} {p' q' : tangent_space I z →L[𝕜] F'} + +lemma has_mfderiv_within_at.mul' (hp : has_mfderiv_within_at I 𝓘(𝕜, F') p s z p') + (hq : has_mfderiv_within_at I 𝓘(𝕜, F') q s z q') : + has_mfderiv_within_at I 𝓘(𝕜, F') (p * q) s z (p z • q' + p'.smul_right (q z) : E →L[𝕜] F') := +⟨hp.1.mul hq.1, by simpa only with mfld_simps using hp.2.mul' hq.2⟩ + +lemma has_mfderiv_at.mul' (hp : has_mfderiv_at I 𝓘(𝕜, F') p z p') + (hq : has_mfderiv_at I 𝓘(𝕜, F') q z q') : + has_mfderiv_at I 𝓘(𝕜, F') (p * q) z (p z • q' + p'.smul_right (q z) : E →L[𝕜] F') := +has_mfderiv_within_at_univ.mp $ hp.has_mfderiv_within_at.mul' hq.has_mfderiv_within_at + +lemma mdifferentiable_within_at.mul (hp : mdifferentiable_within_at I 𝓘(𝕜, F') p s z) + (hq : mdifferentiable_within_at I 𝓘(𝕜, F') q s z) : + mdifferentiable_within_at I 𝓘(𝕜, F') (p * q) s z := +(hp.has_mfderiv_within_at.mul' hq.has_mfderiv_within_at).mdifferentiable_within_at + +lemma mdifferentiable_at.mul (hp : mdifferentiable_at I 𝓘(𝕜, F') p z) + (hq : mdifferentiable_at I 𝓘(𝕜, F') q z) : mdifferentiable_at I 𝓘(𝕜, F') (p * q) z := +(hp.has_mfderiv_at.mul' hq.has_mfderiv_at).mdifferentiable_at + +lemma mdifferentiable_on.mul (hp : mdifferentiable_on I 𝓘(𝕜, F') p s) + (hq : mdifferentiable_on I 𝓘(𝕜, F') q s) : mdifferentiable_on I 𝓘(𝕜, F') (p * q) s := +λ x hx, (hp x hx).mul $ hq x hx + +lemma mdifferentiable.mul (hp : mdifferentiable I 𝓘(𝕜, F') p) + (hq : mdifferentiable I 𝓘(𝕜, F') q) : mdifferentiable I 𝓘(𝕜, F') (p * q) := +λ x, (hp x).mul (hq x) + +end algebra_over_ring + +section algebra_over_comm_ring +variables {I} {z : M} {F' : Type*} [normed_comm_ring F'] [normed_algebra 𝕜 F'] + {p q : M → F'} {p' q' : tangent_space I z →L[𝕜] F'} + +lemma has_mfderiv_within_at.mul (hp : has_mfderiv_within_at I 𝓘(𝕜, F') p s z p') + (hq : has_mfderiv_within_at I 𝓘(𝕜, F') q s z q') : + has_mfderiv_within_at I 𝓘(𝕜, F') (p * q) s z (p z • q' + q z • p' : E →L[𝕜] F') := +by { convert hp.mul' hq, ext z, apply mul_comm } + +lemma has_mfderiv_at.mul (hp : has_mfderiv_at I 𝓘(𝕜, F') p z p') + (hq : has_mfderiv_at I 𝓘(𝕜, F') q z q') : + has_mfderiv_at I 𝓘(𝕜, F') (p * q) z (p z • q' + q z • p' : E →L[𝕜] F') := +has_mfderiv_within_at_univ.mp $ hp.has_mfderiv_within_at.mul hq.has_mfderiv_within_at + +end algebra_over_comm_ring + end arithmetic namespace model_with_corners From 966e0cf0685c9cedf8a3283ac69eef4d5f2eaca2 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Wed, 25 Jan 2023 20:47:30 +0000 Subject: [PATCH 0359/1291] refactor(number_theory/primorial): split the proof, golf it (#18238) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit * Add `nat.prime.dvd_choose`, delete less general `dvd_choose_of_middling_prime`. * Add `primorial_pos`, `primorial_add`, `primorial_add_dvd`, and `primorial_add_le`. * Golf some proofs. Here is a proof of `dvd_choose_of_middling_prime` based on `nat.prime.dvd_choose`: ```lean open nat lemma dvd_choose_of_middling_prime (p : ℕ) (is_prime : nat.prime p) (m : ℕ) (p_big : m + 1 < p) (p_small : p ≤ 2 * m + 1) : p ∣ choose (2 * m + 1) (m + 1) := is_prime.dvd_choose p_big (by simpa [two_mul] using p_big.le) p_small ``` Forward-ported in leanprover-community/mathlib4#1770 --- src/algebra/char_p/basic.lean | 3 +- src/data/nat/choose/dvd.lean | 31 ++-- src/number_theory/primorial.lean | 135 ++++++------------ .../polynomial/eisenstein/is_integral.lean | 7 +- 4 files changed, 60 insertions(+), 116 deletions(-) diff --git a/src/algebra/char_p/basic.lean b/src/algebra/char_p/basic.lean index de1806499b205..600f16288cfe3 100644 --- a/src/algebra/char_p/basic.lean +++ b/src/algebra/char_p/basic.lean @@ -144,8 +144,7 @@ begin { intros b h1 h2, suffices : (p.choose b : R) = 0, { rw this, simp }, rw char_p.cast_eq_zero_iff R p, - refine nat.prime.dvd_choose_self (pos_iff_ne_zero.mpr h2) _ (fact.out _), - rwa ← finset.mem_range }, + exact nat.prime.dvd_choose_self (fact.out _) h2 (finset.mem_range.1 h1) }, { intro h1, contrapose! h1, rw finset.mem_range, diff --git a/src/data/nat/choose/dvd.lean b/src/data/nat/choose/dvd.lean index 11c3114090e97..3adae6d9f24aa 100644 --- a/src/data/nat/choose/dvd.lean +++ b/src/data/nat/choose/dvd.lean @@ -19,26 +19,23 @@ open_locale nat namespace prime -lemma dvd_choose_add {p a b : ℕ} (hap : a < p) (hbp : b < p) (h : p ≤ a + b) - (hp : prime p) : p ∣ choose (a + b) a := -have h₁ : p ∣ (a + b)!, from hp.dvd_factorial.2 h, -have h₂ : ¬p ∣ a!, from mt hp.dvd_factorial.1 (not_le_of_gt hap), -have h₃ : ¬p ∣ b!, from mt hp.dvd_factorial.1 (not_le_of_gt hbp), -by - rw [← choose_mul_factorial_mul_factorial (le.intro rfl), mul_assoc, hp.dvd_mul, hp.dvd_mul, - add_tsub_cancel_left a b] at h₁; - exact h₁.resolve_right (not_or_distrib.2 ⟨h₂, h₃⟩) - -lemma dvd_choose_self {p k : ℕ} (hk : 0 < k) (hkp : k < p) (hp : prime p) : - p ∣ choose p k := +lemma dvd_choose_add {p a b : ℕ} (hp : prime p) (hap : a < p) (hbp : b < p) (h : p ≤ a + b) : + p ∣ choose (a + b) a := begin - have r : k + (p - k) = p, - by rw [← add_tsub_assoc_of_le (nat.le_of_lt hkp) k, add_tsub_cancel_left], - have e : p ∣ choose (k + (p - k)) k, - by exact dvd_choose_add hkp (nat.sub_lt (hk.trans hkp) hk) (by rw r) hp, - rwa r at e, + have h₁ : p ∣ (a + b)!, from hp.dvd_factorial.2 h, + rw [← add_choose_mul_factorial_mul_factorial, ← choose_symm_add, hp.dvd_mul, hp.dvd_mul, + hp.dvd_factorial, hp.dvd_factorial] at h₁, + exact (h₁.resolve_right hbp.not_le).resolve_right hap.not_le end +lemma dvd_choose {p a b : ℕ} (hp : prime p) (ha : a < p) (hab : b - a < p) (h : p ≤ b) : + p ∣ choose b a := +have a + (b - a) = b := nat.add_sub_of_le (ha.le.trans h), +this ▸ hp.dvd_choose_add ha hab (this.symm ▸ h) + +lemma dvd_choose_self {p k : ℕ} (hp : prime p) (hk : k ≠ 0) (hkp : k < p) : p ∣ choose p k := +hp.dvd_choose hkp (nat.sub_lt ((zero_le _).trans_lt hkp) hk.bot_lt) le_rfl + end prime end nat diff --git a/src/number_theory/primorial.lean b/src/number_theory/primorial.lean index 9e8a6174d748d..f216faaed13a3 100644 --- a/src/number_theory/primorial.lean +++ b/src/number_theory/primorial.lean @@ -1,13 +1,13 @@ /- Copyright (c) 2020 Patrick Stevens. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Patrick Stevens +Authors: Patrick Stevens, Yury Kudryashov -/ import algebra.big_operators.associated import data.nat.choose.sum +import data.nat.choose.dvd import data.nat.parity import data.nat.prime -import tactic.ring_exp /-! # Primorial @@ -30,105 +30,54 @@ open_locale big_operators nat def primorial (n : ℕ) : ℕ := ∏ p in (filter nat.prime (range (n + 1))), p local notation x`#` := primorial x -lemma primorial_succ {n : ℕ} (n_big : 1 < n) (r : n % 2 = 1) : (n + 1)# = n# := +lemma primorial_pos (n : ℕ) : 0 < n# := +prod_pos $ λ p hp, (mem_filter.1 hp).2.pos + +lemma primorial_succ {n : ℕ} (hn1 : n ≠ 1) (hn : odd n) : (n + 1)# = n# := begin refine prod_congr _ (λ _ _, rfl), - rw [range_succ, filter_insert, if_neg (λ h, _)], - have two_dvd : 2 ∣ n + 1 := dvd_iff_mod_eq_zero.mpr (by rw [← mod_add_mod, r, mod_self]), - linarith [(h.dvd_iff_eq (nat.bit0_ne_one 1)).mp two_dvd], + rw [range_succ, filter_insert, if_neg (λ h, odd_iff_not_even.mp hn _)], + exact (h.even_sub_one $ mt succ.inj hn1) end -lemma dvd_choose_of_middling_prime (p : ℕ) (is_prime : nat.prime p) (m : ℕ) - (p_big : m + 1 < p) (p_small : p ≤ 2 * m + 1) : p ∣ choose (2 * m + 1) (m + 1) := +lemma primorial_add (m n : ℕ) : + (m + n)# = m# * ∏ p in filter nat.prime (Ico (m + 1) (m + n + 1)), p := begin - have m_size : m + 1 ≤ 2 * m + 1 := le_of_lt (lt_of_lt_of_le p_big p_small), - have s : ¬(p ∣ (m + 1)!), - { intros p_div_fact, - exact lt_le_antisymm p_big (is_prime.dvd_factorial.mp p_div_fact), }, - have t : ¬(p ∣ (2 * m + 1 - (m + 1))!), - { intros p_div_fact, - refine lt_le_antisymm (lt_of_succ_lt p_big) _, - convert is_prime.dvd_factorial.mp p_div_fact, - rw [two_mul, add_assoc, nat.add_sub_cancel] }, - have expanded : - choose (2 * m + 1) (m + 1) * (m + 1)! * (2 * m + 1 - (m + 1))! = (2 * m + 1)! := - @choose_mul_factorial_mul_factorial (2 * m + 1) (m + 1) m_size, - have p_div_big_fact : p ∣ (2 * m + 1)! := (prime.dvd_factorial is_prime).mpr p_small, - rw [←expanded, mul_assoc] at p_div_big_fact, - obtain p_div_choose | p_div_facts : p ∣ choose (2 * m + 1) (m + 1) ∨ p ∣ _! * _! := - (prime.dvd_mul is_prime).1 p_div_big_fact, - { exact p_div_choose, }, - cases (prime.dvd_mul is_prime).1 p_div_facts, - exacts [(s h).elim, (t h).elim], + rw [primorial, primorial, ← Ico_zero_eq_range, ← prod_union, ← filter_union, + Ico_union_Ico_eq_Ico], + exacts [zero_le _, add_le_add_right (nat.le_add_right _ _) _, + disjoint_filter_filter $ Ico_disjoint_Ico_consecutive _ _ _] end -lemma primorial_le_4_pow : ∀ (n : ℕ), n# ≤ 4 ^ n -| 0 := le_rfl -| 1 := le_of_inf_eq rfl -| (n + 2) := - match nat.mod_two_eq_zero_or_one (n + 1) with - | or.inl n_odd := - match nat.even_iff.2 n_odd with - | ⟨m, twice_m⟩ := - have recurse : m + 1 < n + 2 := by linarith, - begin - calc (n + 2)# - = ∏ i in filter nat.prime (range (2 * m + 2)), i : by simpa [two_mul, ←twice_m] - ... = ∏ i in filter nat.prime (finset.Ico (m + 2) (2 * m + 2) ∪ range (m + 2)), i : - begin - rw [range_eq_Ico, finset.union_comm, finset.Ico_union_Ico_eq_Ico], - { exact bot_le }, - { simpa only [add_le_add_iff_right, two_mul] using nat.le_add_left m m }, - end - ... = ∏ i in (filter nat.prime (finset.Ico (m + 2) (2 * m + 2)) - ∪ (filter nat.prime (range (m + 2)))), i : - by rw filter_union - ... = (∏ i in filter nat.prime (finset.Ico (m + 2) (2 * m + 2)), i) - * (∏ i in filter nat.prime (range (m + 2)), i) : - begin - apply finset.prod_union, - have disj : disjoint (finset.Ico (m + 2) (2 * m + 2)) (range (m + 2)), - { simp only [finset.disjoint_left, and_imp, finset.mem_Ico, not_lt, - finset.mem_range], - intros _ pr _, exact pr, }, - exact finset.disjoint_filter_filter disj, - end - ... ≤ (∏ i in filter nat.prime (finset.Ico (m + 2) (2 * m + 2)), i) * 4 ^ (m + 1) : - nat.mul_le_mul_left _ (primorial_le_4_pow (m + 1)) - ... ≤ (choose (2 * m + 1) (m + 1)) * 4 ^ (m + 1) : - begin - have s : ∏ i in filter nat.prime (finset.Ico (m + 2) (2 * m + 2)), - i ∣ choose (2 * m + 1) (m + 1), - { refine prod_primes_dvd (choose (2 * m + 1) (m + 1)) _ _, - { intros a, rw [finset.mem_filter, nat.prime_iff], apply and.right, }, - { intros a, rw finset.mem_filter, - intros pr, - rcases pr with ⟨ size, is_prime ⟩, - simp only [finset.mem_Ico] at size, - rcases size with ⟨ a_big , a_small ⟩, - exact dvd_choose_of_middling_prime a is_prime m a_big - (nat.lt_succ_iff.mp a_small), }, }, - have r : ∏ i in filter nat.prime (finset.Ico (m + 2) (2 * m + 2)), - i ≤ choose (2 * m + 1) (m + 1), - { refine @nat.le_of_dvd _ _ _ s, - exact @choose_pos (2 * m + 1) (m + 1) (by linarith), }, - exact nat.mul_le_mul_right _ r, - end - ... = (choose (2 * m + 1) m) * 4 ^ (m + 1) : by rw choose_symm_half m - ... ≤ 4 ^ m * 4 ^ (m + 1) : nat.mul_le_mul_right _ (choose_middle_le_pow m) - ... = 4 ^ (2 * m + 1) : by ring_exp - ... = 4 ^ (n + 2) : by rw [two_mul, ←twice_m], - end - end - | or.inr n_even := +lemma primorial_add_dvd {m n : ℕ} (h : n ≤ m) : (m + n)# ∣ m# * choose (m + n) m := +calc (m + n)# = m# * ∏ p in filter nat.prime (Ico (m + 1) (m + n + 1)), p : + primorial_add _ _ +... ∣ m# * choose (m + n) m : + mul_dvd_mul_left _ $ prod_primes_dvd _ (λ k hk, (mem_filter.1 hk).2.prime) $ λ p hp, begin - obtain one_lt_n | n_le_one : 1 < n + 1 ∨ n + 1 ≤ 1 := lt_or_le 1 (n + 1), - { rw primorial_succ one_lt_n n_even, - calc (n + 1)# - ≤ 4 ^ n.succ : primorial_le_4_pow (n + 1) - ... ≤ 4 ^ (n + 2) : pow_le_pow (by norm_num) (nat.le_succ _), }, - { have n_zero : n = 0 := eq_bot_iff.2 (succ_le_succ_iff.1 n_le_one), - norm_num [n_zero, primorial, range_succ, prod_filter, nat.not_prime_zero, nat.prime_two] }, + rw [mem_filter, mem_Ico] at hp, + exact hp.2.dvd_choose_add hp.1.1 (h.trans_lt (m.lt_succ_self.trans_le hp.1.1)) + (nat.lt_succ_iff.1 hp.1.2) end +lemma primorial_add_le {m n : ℕ} (h : n ≤ m) : (m + n)# ≤ m# * choose (m + n) m := +le_of_dvd (mul_pos (primorial_pos _) (choose_pos $ nat.le_add_right _ _)) (primorial_add_dvd h) + +lemma primorial_le_4_pow (n : ℕ) : n# ≤ 4 ^ n := +begin + induction n using nat.strong_induction_on with n ihn, + cases n, { refl }, + rcases n.even_or_odd with (⟨m, rfl⟩ | ho), + { rcases m.eq_zero_or_pos with rfl | hm, { dec_trivial }, + calc (m + m + 1)# = (m + 1 + m)# : by rw [add_right_comm] + ... ≤ (m + 1)# * choose (m + 1 + m) (m + 1) : primorial_add_le m.le_succ + ... = (m + 1)# * choose (2 * m + 1) m : by rw [choose_symm_add, two_mul, add_right_comm] + ... ≤ 4 ^ (m + 1) * 4 ^ m : + mul_le_mul' (ihn _ $ succ_lt_succ $ (lt_add_iff_pos_left _).2 hm) (choose_middle_le_pow _) + ... ≤ 4 ^ (m + m + 1) : by rw [← pow_add, add_right_comm] }, + { rcases decidable.eq_or_ne n 1 with rfl | hn, + { dec_trivial }, + { calc (n + 1)# = n# : primorial_succ hn ho + ... ≤ 4 ^ n : ihn n n.lt_succ_self + ... ≤ 4 ^ (n + 1) : pow_le_pow_of_le_right four_pos n.le_succ } } end diff --git a/src/ring_theory/polynomial/eisenstein/is_integral.lean b/src/ring_theory/polynomial/eisenstein/is_integral.lean index 9287f49e81052..79c85cb016f1a 100644 --- a/src/ring_theory/polynomial/eisenstein/is_integral.lean +++ b/src/ring_theory/polynomial/eisenstein/is_integral.lean @@ -54,10 +54,9 @@ begin rw [lcoeff_apply, ← C_eq_nat_cast, C_mul_X_pow_eq_monomial, coeff_monomial] }, rw [nat_degree_comp, show (X + 1 : ℤ[X]) = X + C 1, by simp, nat_degree_X_add_C, mul_one, nat_degree_cyclotomic, nat.totient_prime hp.out] at hi, - simp only [lt_of_lt_of_le hi (nat.sub_le _ _), sum_ite_eq', mem_range, - if_true, ideal.submodule_span_eq, ideal.mem_span_singleton], - exact int.coe_nat_dvd.2 - (nat.prime.dvd_choose_self (nat.succ_pos i) (lt_tsub_iff_right.1 hi) hp.out) }, + simp only [hi.trans_le (nat.sub_le _ _), sum_ite_eq', mem_range, if_true, + ideal.submodule_span_eq, ideal.mem_span_singleton, int.coe_nat_dvd], + exact hp.out.dvd_choose_self i.succ_ne_zero (lt_tsub_iff_right.1 hi) }, { rw [coeff_zero_eq_eval_zero, eval_comp, cyclotomic_prime, eval_add, eval_X, eval_one, zero_add, eval_geom_sum, one_geom_sum, ideal.submodule_span_eq, ideal.span_singleton_pow, ideal.mem_span_singleton], From 0ee3e6f52678a420c58b1072870fed5e0240c083 Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Thu, 26 Jan 2023 08:02:55 +0000 Subject: [PATCH 0360/1291] feat(analysis/special_functions/gamma): monotonicity of the Gamma function (#18246) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Show (using the log-convexity results proved in #18188) that Gamma is strictly increasing for `2 ≤ x`. --- src/analysis/convex/function.lean | 22 +++--- src/analysis/convex/slope.lean | 21 +++++ src/analysis/special_functions/gamma.lean | 82 +++++++++++++++++++- src/analysis/special_functions/gaussian.lean | 37 ++++----- 4 files changed, 127 insertions(+), 35 deletions(-) diff --git a/src/analysis/convex/function.lean b/src/analysis/convex/function.lean index 27795ab9b7fed..920791e909dd6 100644 --- a/src/analysis/convex/function.lean +++ b/src/analysis/convex/function.lean @@ -660,29 +660,25 @@ variables [module 𝕜 E] [module 𝕜 β] [ordered_smul 𝕜 β] {s : set E} {f /- The following lemmas don't require `module 𝕜 E` if you add the hypothesis `x ≠ y`. At the time of the writing, we decided the resulting lemmas wouldn't be useful. Feel free to reintroduce them. -/ -lemma strict_convex_on.lt_left_of_right_lt' (hf : strict_convex_on 𝕜 s f) {x y : E} (hx : x ∈ s) +lemma convex_on.lt_left_of_right_lt' (hf : convex_on 𝕜 s f) {x y : E} (hx : x ∈ s) (hy : y ∈ s) {a b : 𝕜} (ha : 0 < a) (hb : 0 < b) (hab : a + b = 1) (hfy : f y < f (a • x + b • y)) : f (a • x + b • y) < f x := not_le.1 $ λ h, lt_irrefl (f (a • x + b • y)) $ calc f (a • x + b • y) - < a • f x + b • f y : hf.2 hx hy begin - rintro rfl, - rw convex.combo_self hab at hfy, - exact lt_irrefl _ hfy, - end ha hb hab + ≤ a • f x + b • f y : hf.2 hx hy ha.le hb.le hab ... < a • f (a • x + b • y) + b • f (a • x + b • y) : add_lt_add_of_le_of_lt (smul_le_smul_of_nonneg h ha.le) (smul_lt_smul_of_pos hfy hb) ... = f (a • x + b • y) : convex.combo_self hab _ -lemma strict_concave_on.left_lt_of_lt_right' (hf : strict_concave_on 𝕜 s f) {x y : E} (hx : x ∈ s) +lemma concave_on.left_lt_of_lt_right' (hf : concave_on 𝕜 s f) {x y : E} (hx : x ∈ s) (hy : y ∈ s) {a b : 𝕜} (ha : 0 < a) (hb : 0 < b) (hab : a + b = 1) (hfy : f (a • x + b • y) < f y) : f x < f (a • x + b • y) := hf.dual.lt_left_of_right_lt' hx hy ha hb hab hfy -lemma strict_convex_on.lt_right_of_left_lt' (hf : strict_convex_on 𝕜 s f) {x y : E} {a b : 𝕜} +lemma convex_on.lt_right_of_left_lt' (hf : convex_on 𝕜 s f) {x y : E} {a b : 𝕜} (hx : x ∈ s) (hy : y ∈ s) (ha : 0 < a) (hb : 0 < b) (hab : a + b = 1) (hfx : f x < f (a • x + b • y)) : f (a • x + b • y) < f y := @@ -691,13 +687,13 @@ begin exact hf.lt_left_of_right_lt' hy hx hb ha hab hfx, end -lemma strict_concave_on.lt_right_of_left_lt' (hf : strict_concave_on 𝕜 s f) {x y : E} {a b : 𝕜} +lemma concave_on.lt_right_of_left_lt' (hf : concave_on 𝕜 s f) {x y : E} {a b : 𝕜} (hx : x ∈ s) (hy : y ∈ s) (ha : 0 < a) (hb : 0 < b) (hab : a + b = 1) (hfx : f (a • x + b • y) < f x) : f y < f (a • x + b • y) := hf.dual.lt_right_of_left_lt' hx hy ha hb hab hfx -lemma strict_convex_on.lt_left_of_right_lt (hf : strict_convex_on 𝕜 s f) {x y z : E} (hx : x ∈ s) +lemma convex_on.lt_left_of_right_lt (hf : convex_on 𝕜 s f) {x y z : E} (hx : x ∈ s) (hy : y ∈ s) (hz : z ∈ open_segment 𝕜 x y) (hyz : f y < f z) : f z < f x := begin @@ -705,12 +701,12 @@ begin exact hf.lt_left_of_right_lt' hx hy ha hb hab hyz, end -lemma strict_concave_on.left_lt_of_lt_right (hf : strict_concave_on 𝕜 s f) {x y z : E} (hx : x ∈ s) +lemma concave_on.left_lt_of_lt_right (hf : concave_on 𝕜 s f) {x y z : E} (hx : x ∈ s) (hy : y ∈ s) (hz : z ∈ open_segment 𝕜 x y) (hyz : f z < f y) : f x < f z := hf.dual.lt_left_of_right_lt hx hy hz hyz -lemma strict_convex_on.lt_right_of_left_lt (hf : strict_convex_on 𝕜 s f) {x y z : E} (hx : x ∈ s) +lemma convex_on.lt_right_of_left_lt (hf : convex_on 𝕜 s f) {x y z : E} (hx : x ∈ s) (hy : y ∈ s) (hz : z ∈ open_segment 𝕜 x y) (hxz : f x < f z) : f z < f y := begin @@ -718,7 +714,7 @@ begin exact hf.lt_right_of_left_lt' hx hy ha hb hab hxz, end -lemma strict_concave_on.lt_right_of_left_lt (hf : strict_concave_on 𝕜 s f) {x y z : E} (hx : x ∈ s) +lemma concave_on.lt_right_of_left_lt (hf : concave_on 𝕜 s f) {x y z : E} (hx : x ∈ s) (hy : y ∈ s) (hz : z ∈ open_segment 𝕜 x y) (hxz : f z < f x) : f y < f z := hf.dual.lt_right_of_left_lt hx hy hz hxz diff --git a/src/analysis/convex/slope.lean b/src/analysis/convex/slope.lean index abe5e37c12a7a..103c2f74bca6f 100644 --- a/src/analysis/convex/slope.lean +++ b/src/analysis/convex/slope.lean @@ -212,3 +212,24 @@ lemma strict_concave_on_iff_slope_strict_anti_adjacent : (f z - f y) / (z - y) < (f y - f x) / (y - x) := ⟨λ h, ⟨h.1, λ x y z, h.slope_anti_adjacent⟩, λ h, strict_concave_on_of_slope_strict_anti_adjacent h.1 h.2⟩ + +/-- If `f` is convex on a set `s` in a linearly ordered field, and `f x < f y` for two points +`x < y` in `s`, then `f` is strictly monotone on `s ∩ [y, ∞)`. -/ +lemma convex_on.strict_mono_of_lt (hf : convex_on 𝕜 s f) + {x y : 𝕜} (hx : x ∈ s) (hxy : x < y) (hxy' : f x < f y) : + strict_mono_on f (s ∩ set.Ici y) := +begin + intros u hu v hv huv, + have step1 : ∀ {z : 𝕜}, z ∈ s ∩ set.Ioi y → f y < f z, + { refine λ z hz, hf.lt_right_of_left_lt hx hz.1 _ hxy', + rw open_segment_eq_Ioo (hxy.trans hz.2), + exact ⟨hxy, hz.2⟩ }, + rcases eq_or_lt_of_le hu.2 with rfl | hu2, + { exact step1 ⟨hv.1, huv⟩ }, + { refine hf.lt_right_of_left_lt _ hv.1 _ (step1 ⟨hu.1, hu2⟩), + { apply hf.1.segment_subset hx hu.1, + rw segment_eq_Icc (hxy.le.trans hu.2), + exact ⟨hxy.le, hu.2⟩ }, + { rw open_segment_eq_Ioo (hu2.trans huv), + exact ⟨hu2, huv⟩ } }, +end diff --git a/src/analysis/special_functions/gamma.lean b/src/analysis/special_functions/gamma.lean index d6775ee17112c..5a56c0228d5c1 100644 --- a/src/analysis/special_functions/gamma.lean +++ b/src/analysis/special_functions/gamma.lean @@ -52,7 +52,7 @@ Gamma noncomputable theory open filter interval_integral set real measure_theory asymptotics -open_locale nat topological_space ennreal big_operators +open_locale nat topological_space ennreal big_operators complex_conjugate lemma integral_exp_neg_Ioi : ∫ (x : ℝ) in Ioi 0, exp (-x) = 1 := begin @@ -132,6 +132,16 @@ See `complex.Gamma_integral_convergent` for a proof of the convergence of the in `0 < re s`. -/ def Gamma_integral (s : ℂ) : ℂ := ∫ x in Ioi (0:ℝ), ↑(-x).exp * ↑x ^ (s - 1) +lemma Gamma_integral_conj (s : ℂ) : Gamma_integral (conj s) = conj (Gamma_integral s) := +begin + rw [Gamma_integral, Gamma_integral, ←integral_conj], + refine set_integral_congr measurable_set_Ioi (λ x hx, _), + dsimp only, + rw [ring_hom.map_mul, conj_of_real, cpow_def_of_ne_zero (of_real_ne_zero.mpr (ne_of_gt hx)), + cpow_def_of_ne_zero (of_real_ne_zero.mpr (ne_of_gt hx)), ←exp_conj, ring_hom.map_mul, + ←of_real_log (le_of_lt hx), conj_of_real, ring_hom.map_sub, ring_hom.map_one], +end + lemma Gamma_integral_of_real (s : ℝ) : Gamma_integral ↑s = ↑(∫ x:ℝ in Ioi 0, real.exp (-x) * x ^ (s - 1)) := begin @@ -307,7 +317,6 @@ begin rw this }, end - /-- The `Γ` function (of a complex variable `s`). -/ @[pp_nodot] def Gamma (s : ℂ) : ℂ := Gamma_aux ⌊1 - s.re⌋₊ s @@ -356,6 +365,20 @@ begin simp only [nat.cast_succ, nat.factorial_succ, nat.cast_mul], congr, exact hn }, end +lemma Gamma_conj (s : ℂ) : Gamma (conj s) = conj (Gamma s) := +begin + suffices : ∀ (n:ℕ) (s:ℂ) , Gamma_aux n (conj s) = conj (Gamma_aux n s), from this _ _, + intro n, + induction n with n IH, + { rw Gamma_aux, exact Gamma_integral_conj, }, + { intro s, + rw Gamma_aux, + dsimp only, + rw [div_eq_mul_inv _ s, ring_hom.map_mul, conj_inv, ←div_eq_mul_inv], + suffices : conj s + 1 = conj (s + 1), by rw [this, IH], + rw [ring_hom.map_add, ring_hom.map_one] } +end + end Gamma_def end complex @@ -566,6 +589,9 @@ end lemma Gamma_one : Gamma 1 = 1 := by rw [Gamma, complex.of_real_one, complex.Gamma_one, complex.one_re] +lemma _root_.complex.Gamma_of_real (s : ℝ) : complex.Gamma (s : ℂ) = Gamma s := +by rw [Gamma, eq_comm, ←complex.eq_conj_iff_re, ←complex.Gamma_conj, complex.conj_of_real] + theorem Gamma_nat_eq_factorial (n : ℕ) : Gamma (n + 1) = n! := by rw [Gamma, complex.of_real_add, complex.of_real_nat_cast, complex.of_real_one, complex.Gamma_nat_eq_factorial, ←complex.of_real_nat_cast, complex.of_real_re] @@ -697,6 +723,21 @@ begin exact Gamma_mul_add_mul_le_rpow_Gamma_mul_rpow_Gamma hx hy ha hb hab, end +lemma convex_on_Gamma : convex_on ℝ (Ioi 0) Gamma := +begin + refine ⟨convex_Ioi 0, λ x hx y hy a b ha hb hab, _⟩, + have := convex_on.comp (convex_on_exp.subset (subset_univ _) _) convex_on_log_Gamma + (λ u hu v hv huv, exp_le_exp.mpr huv), + convert this.2 hx hy ha hb hab, + { rw [function.comp_app, exp_log (Gamma_pos_of_pos $ this.1 hx hy ha hb hab)] }, + { rw [function.comp_app, exp_log (Gamma_pos_of_pos hx)] }, + { rw [function.comp_app, exp_log (Gamma_pos_of_pos hy)] }, + { rw convex_iff_is_preconnected, + refine is_preconnected_Ioi.image _ (λ x hx, continuous_at.continuous_within_at _), + refine (differentiable_at_Gamma (λ m, _)).continuous_at.log (Gamma_pos_of_pos hx).ne', + exact (add_pos_of_pos_of_nonneg hx (nat.cast_nonneg m)).ne' }, +end + section bohr_mollerup /-! ## The Euler limit formula and the Bohr-Mollerup theorem @@ -928,4 +969,41 @@ end end bohr_mollerup +section strict_mono + +lemma Gamma_two : Gamma 2 = 1 := by simpa using Gamma_nat_eq_factorial 1 + +lemma Gamma_three_div_two_lt_one : Gamma (3 / 2) < 1 := +begin + -- This can also be proved using the closed-form evaluation of `Gamma (1 / 2)` in + -- `analysis.special_functions.gaussian`, but we give a self-contained proof using log-convexity + -- to avoid unnecessary imports. + have A : (0:ℝ) < 3/2, by norm_num, + have := bohr_mollerup.f_add_nat_le convex_on_log_Gamma (λ y hy, _) two_ne_zero one_half_pos + (by norm_num : 1/2 ≤ (1:ℝ)), + swap, { rw [function.comp_app, Gamma_add_one hy.ne', log_mul hy.ne' (Gamma_pos_of_pos hy).ne', + add_comm] }, + rw [function.comp_app, function.comp_app, nat.cast_two, Gamma_two, log_one, zero_add, + (by norm_num : (2:ℝ) + 1/2 = 3/2 + 1), Gamma_add_one A.ne', + log_mul A.ne' (Gamma_pos_of_pos A).ne', ←le_sub_iff_add_le', + log_le_iff_le_exp (Gamma_pos_of_pos A)] at this, + refine this.trans_lt (exp_lt_one_iff.mpr _), + rw [mul_comm, ←mul_div_assoc, div_sub' _ _ (2:ℝ) two_ne_zero], + refine div_neg_of_neg_of_pos _ two_pos, + rw [sub_neg, mul_one, ←nat.cast_two, ←log_pow, ←exp_lt_exp, nat.cast_two, exp_log two_pos, + exp_log]; + norm_num, +end + +lemma Gamma_strict_mono_on_Ici : strict_mono_on Gamma (Ici 2) := +begin + convert convex_on_Gamma.strict_mono_of_lt (by norm_num : (0:ℝ) < 3/2) + (by norm_num : (3/2 : ℝ) < 2) (Gamma_two.symm ▸ Gamma_three_div_two_lt_one), + symmetry, + rw inter_eq_right_iff_subset, + exact λ x hx, two_pos.trans_le hx, +end + +end strict_mono + end real diff --git a/src/analysis/special_functions/gaussian.lean b/src/analysis/special_functions/gaussian.lean index 8053436bc7b1a..9eabecd37a2f4 100644 --- a/src/analysis/special_functions/gaussian.lean +++ b/src/analysis/special_functions/gaussian.lean @@ -309,31 +309,28 @@ begin exact (div_pos pi_pos hb).le, } end -namespace complex - /-- The special-value formula `Γ(1/2) = √π`, which is equivalent to the Gaussian integral. -/ -lemma Gamma_one_half_eq : Gamma (1 / 2) = sqrt π := +lemma real.Gamma_one_half_eq : real.Gamma (1 / 2) = sqrt π := begin - -- first reduce to real integrals - have hh : (1 / 2 : ℂ) = ↑(1 / 2 : ℝ), - { simp only [one_div, of_real_inv, of_real_bit0, of_real_one] }, - have hh2 : (1 / 2 : ℂ).re = 1 / 2, - { convert of_real_re (1 / 2 : ℝ) }, - replace hh2 : 0 < (1 / 2 : ℂ).re := by { rw hh2, exact one_half_pos, }, - rw [Gamma_eq_integral hh2, hh, Gamma_integral_of_real, of_real_inj], - -- now do change-of-variables - rw ←integral_comp_rpow_Ioi_of_pos zero_lt_two, - have : eq_on (λ x:ℝ, (2 * x^((2:ℝ) - 1)) • (real.exp (-x^(2:ℝ)) * (x^(2:ℝ)) ^ (1 / (2:ℝ) - 1))) - (λ x:ℝ, 2 * real.exp ((-1) * x ^ (2:ℕ))) (Ioi 0), - { intros x hx, dsimp only, - have : (x^(2:ℝ)) ^ (1 / (2:ℝ) - 1) = x⁻¹, - { rw ←rpow_mul (le_of_lt hx), norm_num, + rw [Gamma_eq_integral one_half_pos, ←integral_comp_rpow_Ioi_of_pos zero_lt_two], + convert congr_arg (λ x:ℝ, 2 * x) (integral_gaussian_Ioi 1), + { rw ←integral_mul_left, + refine set_integral_congr measurable_set_Ioi (λ x hx, _), + dsimp only, + have : (x ^ (2:ℝ)) ^ (1 / (2:ℝ) - 1) = x⁻¹, + { rw ←rpow_mul (le_of_lt hx), + norm_num, rw [rpow_neg (le_of_lt hx), rpow_one] }, rw [smul_eq_mul, this], field_simp [(ne_of_lt hx).symm], norm_num, ring }, - rw [set_integral_congr measurable_set_Ioi this, integral_mul_left, integral_gaussian_Ioi], - field_simp, ring, + { rw [div_one, ←mul_div_assoc, mul_comm, mul_div_cancel _ (two_ne_zero' ℝ)], } end -end complex +/-- The special-value formula `Γ(1/2) = √π`, which is equivalent to the Gaussian integral. -/ +lemma complex.Gamma_one_half_eq : complex.Gamma (1 / 2) = π ^ (1 / 2 : ℂ) := +begin + convert congr_arg coe real.Gamma_one_half_eq, + { simpa only [one_div, of_real_inv, of_real_bit0] using Gamma_of_real (1 / 2)}, + { rw [sqrt_eq_rpow, of_real_cpow pi_pos.le, of_real_div, of_real_bit0, of_real_one] } +end From 397a33d967c1f683a2eb1e530672d6b5818d6536 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Thu, 26 Jan 2023 09:45:55 +0000 Subject: [PATCH 0361/1291] chore(set_theory/game/nim): golf `grundy_value_nim_add_nim` (#15878) We also remove `exists_ordinal_move_left_eq` and `exists_move_left_eq`, as they just express in a roundabout way that `to_left_moves_nim` is an equivalence. --- src/set_theory/game/nim.lean | 93 +++++++++----------------- src/set_theory/ordinal/arithmetic.lean | 8 +++ 2 files changed, 38 insertions(+), 63 deletions(-) diff --git a/src/set_theory/game/nim.lean b/src/set_theory/game/nim.lean index 64a8e2d305989..d768c21db2dee 100644 --- a/src/set_theory/game/nim.lean +++ b/src/set_theory/game/nim.lean @@ -182,12 +182,6 @@ begin simpa using IH _ (typein_lt_self _) end -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 o' : ordinal} (h : o' < o) : ∃ i, (nim o).move_left i = nim o' := -⟨to_left_moves_nim ⟨o', h⟩, by simp⟩ - lemma nim_fuzzy_zero_of_ne_zero {o : ordinal} (ho : o ≠ 0) : nim o ‖ 0 := begin rw [impartial.fuzzy_zero_iff_lf, nim_def, lf_zero_le], @@ -303,73 +297,46 @@ lemma grundy_value_eq_mex_right : ∀ (G : pgame) [G.impartial], apply grundy_value_neg end +/-- The Grundy value of the sum of two nim games with natural numbers of piles equals their bitwise +xor. -/ +-- Todo: this actually generalizes to all ordinals, by defining `ordinal.lxor` as the pairwise +-- `nat.lxor` of base `ω` Cantor normal forms. @[simp] lemma grundy_value_nim_add_nim (n m : ℕ) : grundy_value (nim.{u} n + nim.{u} m) = nat.lxor n m := begin + -- We do strong induction on both variables. induction n using nat.strong_induction_on with n hn generalizing m, induction m using nat.strong_induction_on with m hm, - rw [grundy_value_eq_mex_left], - - -- We want to show that `n xor m` is the smallest unreachable Grundy value. We will do this in two - -- steps: - -- h₀: `n xor m` is not a reachable grundy number. - -- h₁: every Grundy number strictly smaller than `n xor m` is reachable. - - have h₀ : ∀ i, grundy_value ((nim n + nim m).move_left i) ≠ (nat.lxor n m : ordinal), - { -- To show that `n xor m` is unreachable, we show that every move produces a Grundy number - -- different from `n xor m`. - intro i, - - -- The move operates either on the left pile or on the right pile. - apply left_moves_add_cases i, - - all_goals - { -- One of the piles is reduced to `k` stones, with `k < n` or `k < m`. - intro a, - obtain ⟨ok, hk, hk'⟩ := 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, - - -- Thus, the problem is reduced to computing the Grundy value of `nim n + nim k` or - -- `nim k + nim m`, both of which can be dealt with using an inductive hypothesis. - simp only [hk', add_move_left_inl, add_move_left_inr, id], + rw grundy_value_eq_mex_left, + apply (ordinal.mex_le_of_ne.{u u} (λ i, _)).antisymm (ordinal.le_mex_of_forall (λ ou hu, _)), + -- The Grundy value `nat.lxor n m` can't be reached by left moves. + { apply left_moves_add_cases i; + { -- A left move leaves us with a Grundy value of `nat.lxor k m` for `k < n`, or `nat.lxor n k` + -- for `k < m`. + refine λ a, left_moves_nim_rec_on a (λ ok hk, _), + obtain ⟨k, rfl⟩ := ordinal.lt_omega.1 (hk.trans (ordinal.nat_lt_omega _)), + simp only [add_move_left_inl, add_move_left_inr, move_left_nim', equiv.symm_apply_apply], + + -- The inequality follows from injectivity. + rw nat_cast_lt at hk, rw hn _ hk <|> rw hm _ hk, - - -- But of course xor is injective, so if we change one of the arguments, we will not get the - -- same value again. - intro h, + refine λ h, hk.ne _, rw ordinal.nat_cast_inj at h, - try { rw [nat.lxor_comm n k, nat.lxor_comm n m] at h }, - exact hk.ne (nat.lxor_left_injective h) } }, - - have h₁ : ∀ (u : ordinal), u < nat.lxor n m → - u ∈ set.range (λ i, grundy_value ((nim n + nim m).move_left i)), - { -- Take any natural number `u` less than `n xor m`. - intros ou hu, - obtain ⟨u, rfl⟩ := ordinal.lt_omega.1 (lt_trans hu (ordinal.nat_lt_omega _)), + rwa nat.lxor_left_inj at h <|> rwa nat.lxor_right_inj at h } }, + -- Every other smaller Grundy value can be reached by left moves. + { -- If `u < nat.lxor m n`, then either `nat.lxor u n < m` or `nat.lxor u m < n`. + obtain ⟨u, rfl⟩ := ordinal.lt_omega.1 (hu.trans (ordinal.nat_lt_omega _)), replace hu := ordinal.nat_cast_lt.1 hu, - - -- Our goal is to produce a move that gives the Grundy value `u`. - rw set.mem_range, - - -- By a lemma about xor, either `u xor m < n` or `u xor n < m`. cases nat.lt_lxor_cases hu with h h, - -- 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⟩ := exists_move_left_eq (ordinal.nat_cast_lt.2 h), - refine ⟨to_left_moves_add (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⟩ := exists_move_left_eq (ordinal.nat_cast_lt.2 h), - refine ⟨to_left_moves_add (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] } }, - - -- We are done! - apply (ordinal.mex_le_of_ne.{u u} h₀).antisymm, - contrapose! h₁, - exact ⟨_, ⟨h₁, ordinal.mex_not_mem_range _⟩⟩, + -- In the first case, reducing the `m` pile to `nat.lxor u n` gives the desired Grundy value. + { refine ⟨to_left_moves_add (sum.inl $ to_left_moves_nim ⟨_, ordinal.nat_cast_lt.2 h⟩), _⟩, + simp [nat.lxor_cancel_right, hn _ h] }, + + -- In the second case, reducing the `n` pile to `nat.lxor u m` gives the desired Grundy value. + { refine ⟨to_left_moves_add (sum.inr $ to_left_moves_nim ⟨_, ordinal.nat_cast_lt.2 h⟩), _⟩, + have : n.lxor (u.lxor n) = u, rw [nat.lxor_comm u, nat.lxor_cancel_left], + simpa [hm _ h] using this } } end lemma nim_add_nim_equiv {n m : ℕ} : nim n + nim m ≈ nim (nat.lxor n m) := diff --git a/src/set_theory/ordinal/arithmetic.lean b/src/set_theory/ordinal/arithmetic.lean index 6057463f2d35b..5f8e3b0edbf3c 100644 --- a/src/set_theory/ordinal/arithmetic.lean +++ b/src/set_theory/ordinal/arithmetic.lean @@ -1512,6 +1512,10 @@ Inf (set.range f)ᶜ theorem mex_not_mem_range {ι : Type u} (f : ι → ordinal.{max u v}) : mex f ∉ set.range f := Inf_mem (nonempty_compl_range f) +theorem le_mex_of_forall {ι : Type u} {f : ι → ordinal.{max u v}} {a : ordinal} + (H : ∀ b < a, ∃ i, f i = b) : a ≤ mex f := +by { by_contra' h, exact mex_not_mem_range f (H _ h) } + theorem ne_mex {ι} (f : ι → ordinal) : ∀ i, f i ≠ mex f := by simpa using mex_not_mem_range f @@ -1559,6 +1563,10 @@ mex (family_of_bfamily o f) theorem bmex_not_mem_brange {o : ordinal} (f : Π a < o, ordinal) : bmex o f ∉ brange o f := by { rw ←range_family_of_bfamily, apply mex_not_mem_range } +theorem le_bmex_of_forall {o : ordinal} (f : Π a < o, ordinal) {a : ordinal} + (H : ∀ b < a, ∃ i hi, f i hi = b) : a ≤ bmex o f := +by { by_contra' h, exact bmex_not_mem_brange f (H _ h) } + theorem ne_bmex {o : ordinal} (f : Π a < o, ordinal) {i} (hi) : f i hi ≠ bmex o f := begin convert ne_mex _ (enum (<) i (by rwa type_lt)), From 8c75ef3517d4106e89fe524e6281d0b0545f47fc Mon Sep 17 00:00:00 2001 From: Amelia Livingston <101damnations@github.com> Date: Thu, 26 Jan 2023 09:45:57 +0000 Subject: [PATCH 0362/1291] feat(algebra/homology/opposite): add opposite complexes (#18144) The opposite of the category of chain complexes with objects in $V$ is equivalent to the category of cochain complexes of objects in $V^{op}.$ Moreover, the opposite of the homology of a chain complex is isomorphic to the cohomology of the corresponding cochain complex of objects in $V^{op}.$ We prove this more generally, for any complex shape. --- src/algebra/homology/opposite.lean | 262 ++++++++++++++++++++++ src/category_theory/abelian/basic.lean | 30 +++ src/category_theory/abelian/opposite.lean | 40 ++++ 3 files changed, 332 insertions(+) create mode 100644 src/algebra/homology/opposite.lean diff --git a/src/algebra/homology/opposite.lean b/src/algebra/homology/opposite.lean new file mode 100644 index 0000000000000..e015eb1f61d5e --- /dev/null +++ b/src/algebra/homology/opposite.lean @@ -0,0 +1,262 @@ +/- +Copyright (c) 2022 Amelia Livingston. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johan Commelin, Amelia Livingston +-/ + +import category_theory.abelian.opposite +import category_theory.abelian.homology +import algebra.homology.additive + +/-! +# Opposite categories of complexes +Given a preadditive category `V`, the opposite of its category of chain complexes is equivalent to +the category of cochain complexes of objects in `Vᵒᵖ`. We define this equivalence, and another +analagous equivalence (for a general category of homological complexes with a general +complex shape). + +We then show that when `V` is abelian, if `C` is a homological complex, then the homology of +`op(C)` is isomorphic to `op` of the homology of `C` (and the analagous result for `unop`). + +## Implementation notes +It is convenient to define both `op` and `op_symm`; this is because given a complex shape `c`, +`c.symm.symm` is not defeq to `c`. + +## Tags +opposite, chain complex, cochain complex, homology, cohomology, homological complex +-/ + + +noncomputable theory + +open opposite category_theory category_theory.limits + +section + +variables {V : Type*} [category V] [abelian V] + +lemma image_to_kernel_op {X Y Z : V} (f : X ⟶ Y) (g : Y ⟶ Z) (w : f ≫ g = 0) : + image_to_kernel g.op f.op (by rw [←op_comp, w, op_zero]) = ((image_subobject_iso _) + ≪≫ (image_op_op _).symm).hom ≫ (cokernel.desc f (factor_thru_image g) + (by rw [←cancel_mono (image.ι g), category.assoc, image.fac, w, zero_comp])).op + ≫ ((kernel_subobject_iso _) ≪≫ (kernel_op_op _)).inv := +begin + ext, + simpa only [iso.trans_hom, iso.symm_hom, iso.trans_inv, kernel_op_op_inv, category.assoc, + image_to_kernel_arrow, kernel_subobject_arrow', kernel.lift_ι, ←op_comp, cokernel.π_desc, + ←image_subobject_arrow, ←image_unop_op_inv_comp_op_factor_thru_image g.op], +end + +lemma image_to_kernel_unop {X Y Z : Vᵒᵖ} (f : X ⟶ Y) (g : Y ⟶ Z) (w : f ≫ g = 0) : + image_to_kernel g.unop f.unop (by rw [←unop_comp, w, unop_zero]) = ((image_subobject_iso _) + ≪≫ (image_unop_unop _).symm).hom ≫ (cokernel.desc f (factor_thru_image g) + (by rw [←cancel_mono (image.ι g), category.assoc, image.fac, w, zero_comp])).unop + ≫ ((kernel_subobject_iso _) ≪≫ (kernel_unop_unop _)).inv := +begin + ext, + dunfold image_unop_unop, + simp only [iso.trans_hom, iso.symm_hom, iso.trans_inv, kernel_unop_unop_inv, category.assoc, + image_to_kernel_arrow, kernel_subobject_arrow', kernel.lift_ι, cokernel.π_desc, + iso.unop_inv, ←unop_comp, factor_thru_image_comp_image_unop_op_inv, quiver.hom.unop_op, + image_subobject_arrow], +end + +/-- Given `f, g` with `f ≫ g = 0`, the homology of `g.op, f.op` is the opposite of the homology of +`f, g`. -/ +def homology_op {X Y Z : V} (f : X ⟶ Y) (g : Y ⟶ Z) (w : f ≫ g = 0) : + homology g.op f.op (by rw [←op_comp, w, op_zero]) ≅ opposite.op (homology f g w) := +cokernel_iso_of_eq (image_to_kernel_op _ _ w) ≪≫ (cokernel_epi_comp _ _) + ≪≫ cokernel_comp_is_iso _ _ ≪≫ cokernel_op_op _ ≪≫ ((homology_iso_kernel_desc _ _ _) + ≪≫ (kernel_iso_of_eq (by ext; simp only [image.fac, cokernel.π_desc, cokernel.π_desc_assoc])) + ≪≫ (kernel_comp_mono _ (image.ι g))).op + +/-- Given morphisms `f, g` in `Vᵒᵖ` with `f ≫ g = 0`, the homology of `g.unop, f.unop` is the +opposite of the homology of `f, g`. -/ +def homology_unop {X Y Z : Vᵒᵖ} (f : X ⟶ Y) (g : Y ⟶ Z) (w : f ≫ g = 0) : + homology g.unop f.unop (by rw [←unop_comp, w, unop_zero]) ≅ opposite.unop (homology f g w) := +cokernel_iso_of_eq (image_to_kernel_unop _ _ w) ≪≫ (cokernel_epi_comp _ _) + ≪≫ cokernel_comp_is_iso _ _ ≪≫ cokernel_unop_unop _ + ≪≫ ((homology_iso_kernel_desc _ _ _) + ≪≫ (kernel_iso_of_eq (by ext; simp only [image.fac, cokernel.π_desc, cokernel.π_desc_assoc])) + ≪≫ (kernel_comp_mono _ (image.ι g))).unop + +end + +namespace homological_complex + +variables {ι V : Type*} [category V] {c : complex_shape ι} + +section +variables [preadditive V] + +/-- Sends a complex `X` with objects in `V` to the corresponding complex with objects in `Vᵒᵖ`. -/ +@[simps] protected def op (X : homological_complex V c) : homological_complex Vᵒᵖ c.symm := +{ X := λ i, op (X.X i), + d := λ i j, (X.d j i).op, + shape' := λ i j hij, by { rw [X.shape j i hij, op_zero], }, + d_comp_d' := by { intros, rw [← op_comp, X.d_comp_d, op_zero], } } + +/-- Sends a complex `X` with objects in `V` to the corresponding complex with objects in `Vᵒᵖ`. -/ +@[simps] protected def op_symm (X : homological_complex V c.symm) : homological_complex Vᵒᵖ c := +{ X := λ i, op (X.X i), + d := λ i j, (X.d j i).op, + shape' := λ i j hij, by { rw [X.shape j i hij, op_zero], }, + d_comp_d' := by { intros, rw [← op_comp, X.d_comp_d, op_zero], } } + +/-- Sends a complex `X` with objects in `Vᵒᵖ` to the corresponding complex with objects in `V`. -/ +@[simps] protected def unop (X : homological_complex Vᵒᵖ c) : homological_complex V c.symm := +{ X := λ i, unop (X.X i), + d := λ i j, (X.d j i).unop, + shape' := λ i j hij, by { rw [X.shape j i hij, unop_zero], }, + d_comp_d' := by { intros, rw [← unop_comp, X.d_comp_d, unop_zero], } } + +/-- Sends a complex `X` with objects in `Vᵒᵖ` to the corresponding complex with objects in `V`. -/ +@[simps] protected def unop_symm (X : homological_complex Vᵒᵖ c.symm) : homological_complex V c := +{ X := λ i, unop (X.X i), + d := λ i j, (X.d j i).unop, + shape' := λ i j hij, by { rw [X.shape j i hij, unop_zero], }, + d_comp_d' := by { intros, rw [← unop_comp, X.d_comp_d, unop_zero], } } + +variables (V c) + +/-- Auxilliary definition for `op_equivalence`. -/ +@[simps] def op_functor : (homological_complex V c)ᵒᵖ ⥤ homological_complex Vᵒᵖ c.symm := +{ obj := λ X, (unop X).op, + map := λ X Y f, + { f := λ i, (f.unop.f i).op, + comm' := λ i j hij, by simp only [op_d, ← op_comp, f.unop.comm] }, } + +/-- Auxilliary definition for `op_equivalence`. -/ +@[simps] def op_inverse : homological_complex Vᵒᵖ c.symm ⥤ (homological_complex V c)ᵒᵖ := +{ obj := λ X, op X.unop_symm, + map := λ X Y f, quiver.hom.op + { f := λ i, (f.f i).unop, + comm' := λ i j hij, by simp only [unop_symm_d, ←unop_comp, f.comm], }} + +/-- Auxilliary definition for `op_equivalence`. -/ +def op_unit_iso : 𝟭 (homological_complex V c)ᵒᵖ ≅ op_functor V c ⋙ op_inverse V c := +nat_iso.of_components (λ X, (homological_complex.hom.iso_of_components (λ i, iso.refl _) + (λ i j hij, by simp only [iso.refl_hom, category.id_comp, unop_symm_d, op_d, quiver.hom.unop_op, + category.comp_id]) : (opposite.unop X).op.unop_symm ≅ unop X).op) + begin + intros X Y f, + refine quiver.hom.unop_inj _, + ext, + simp only [quiver.hom.unop_op, functor.id_map, iso.op_hom, functor.comp_map, + unop_comp, comp_f, hom.iso_of_components_hom_f], + erw [category.id_comp, category.comp_id (f.unop.f x)], + end + +/-- Auxilliary definition for `op_equivalence`. -/ +def op_counit_iso : op_inverse V c ⋙ op_functor V c ≅ 𝟭 (homological_complex Vᵒᵖ c.symm) := +nat_iso.of_components (λ X, homological_complex.hom.iso_of_components (λ i, iso.refl _) + (λ i j hij, by simpa only [iso.refl_hom, category.id_comp, category.comp_id])) + begin + intros X Y f, + ext, + simpa only [quiver.hom.unop_op, quiver.hom.op_unop, functor.comp_map, functor.id_map, + iso.refl_hom, category.id_comp, category.comp_id, comp_f, hom.iso_of_components_hom_f], + end + +/-- Given a category of complexes with objects in `V`, there is a natural equivalence between its +opposite category and a category of complexes with objects in `Vᵒᵖ`. -/ +@[simps] def op_equivalence : (homological_complex V c)ᵒᵖ ≌ homological_complex Vᵒᵖ c.symm := +{ functor := op_functor V c, + inverse := op_inverse V c, + unit_iso := op_unit_iso V c, + counit_iso := op_counit_iso V c, + functor_unit_iso_comp' := + begin + intro X, + ext, + simp only [op_unit_iso, op_counit_iso, nat_iso.of_components_hom_app, iso.op_hom, + comp_f, op_functor_map_f, quiver.hom.unop_op, hom.iso_of_components_hom_f], + exact category.comp_id _, + end } + +/-- Auxilliary definition for `unop_equivalence`. -/ +@[simps] def unop_functor : (homological_complex Vᵒᵖ c)ᵒᵖ ⥤ homological_complex V c.symm := +{ obj := λ X, (unop X).unop, + map := λ X Y f, + { f := λ i, (f.unop.f i).unop, + comm' := λ i j hij, by simp only [unop_d, ← unop_comp, f.unop.comm] }, } + +/-- Auxilliary definition for `unop_equivalence`. -/ +@[simps] def unop_inverse : homological_complex V c.symm ⥤ (homological_complex Vᵒᵖ c)ᵒᵖ := +{ obj := λ X, op X.op_symm, + map := λ X Y f, quiver.hom.op + { f := λ i, (f.f i).op, + comm' := λ i j hij, by simp only [op_symm_d, ←op_comp, f.comm], }} + +/-- Auxilliary definition for `unop_equivalence`. -/ +def unop_unit_iso : 𝟭 (homological_complex Vᵒᵖ c)ᵒᵖ ≅ unop_functor V c ⋙ unop_inverse V c := +nat_iso.of_components (λ X, (homological_complex.hom.iso_of_components (λ i, iso.refl _) + (λ i j hij, by simp only [iso.refl_hom, category.id_comp, unop_symm_d, op_d, quiver.hom.unop_op, + category.comp_id]) : (opposite.unop X).op.unop_symm ≅ unop X).op) + begin + intros X Y f, + refine quiver.hom.unop_inj _, + ext, + simp only [quiver.hom.unop_op, functor.id_map, iso.op_hom, functor.comp_map, + unop_comp, comp_f, hom.iso_of_components_hom_f], + erw [category.id_comp, category.comp_id (f.unop.f x)], + end + +/-- Auxilliary definition for `unop_equivalence`. -/ +def unop_counit_iso : unop_inverse V c ⋙ unop_functor V c ≅ 𝟭 (homological_complex V c.symm) := +nat_iso.of_components (λ X, homological_complex.hom.iso_of_components (λ i, iso.refl _) + (λ i j hij, by simpa only [iso.refl_hom, category.id_comp, category.comp_id])) + begin + intros X Y f, + ext, + simpa only [quiver.hom.unop_op, quiver.hom.op_unop, functor.comp_map, functor.id_map, + iso.refl_hom, category.id_comp, category.comp_id, comp_f, hom.iso_of_components_hom_f], + end + +/-- Given a category of complexes with objects in `Vᵒᵖ`, there is a natural equivalence between its +opposite category and a category of complexes with objects in `V`. -/ +@[simps] def unop_equivalence : (homological_complex Vᵒᵖ c)ᵒᵖ ≌ homological_complex V c.symm := +{ functor := unop_functor V c, + inverse := unop_inverse V c, + unit_iso := unop_unit_iso V c, + counit_iso := unop_counit_iso V c, + functor_unit_iso_comp' := + begin + intro X, + ext, + simp only [op_unit_iso, op_counit_iso, nat_iso.of_components_hom_app, iso.op_hom, + comp_f, op_functor_map_f, quiver.hom.unop_op, hom.iso_of_components_hom_f], + exact category.comp_id _, + end } + +variables {V c} +instance op_functor_additive : (@op_functor ι V _ c _).additive := {} + +instance unop_functor_additive : (@unop_functor ι V _ c _).additive := {} + +end + +variables [abelian V] (C : homological_complex V c) (i : ι) + +/-- Auxilliary tautological definition for `homology_op`. -/ +def homology_op_def : + C.op.homology i ≅ _root_.homology (C.d_from i).op (C.d_to i).op + (by rw [←op_comp, C.d_to_comp_d_from i, op_zero]) := iso.refl _ + +/-- Given a complex `C` of objects in `V`, the `i`th homology of its 'opposite' complex (with +objects in `Vᵒᵖ`) is the opposite of the `i`th homology of `C`. -/ +def homology_op : C.op.homology i ≅ opposite.op (C.homology i) := +homology_op_def _ _ ≪≫ homology_op _ _ _ + +/-- Auxilliary tautological definition for `homology_unop`. -/ +def homology_unop_def (C : homological_complex Vᵒᵖ c) : + C.unop.homology i ≅ _root_.homology (C.d_from i).unop (C.d_to i).unop + (by rw [←unop_comp, C.d_to_comp_d_from i, unop_zero]) := iso.refl _ + +/-- Given a complex `C` of objects in `Vᵒᵖ`, the `i`th homology of its 'opposite' complex (with +objects in `V`) is the opposite of the `i`th homology of `C`. -/ +def homology_unop (C : homological_complex Vᵒᵖ c) : + C.unop.homology i ≅ opposite.unop (C.homology i) := +homology_unop_def _ _ ≪≫ homology_unop _ _ _ + +end homological_complex diff --git a/src/category_theory/abelian/basic.lean b/src/category_theory/abelian/basic.lean index 50595880de1fc..160b7c3b88b08 100644 --- a/src/category_theory/abelian/basic.lean +++ b/src/category_theory/abelian/basic.lean @@ -381,11 +381,41 @@ abbreviation coimage_iso_image' : abelian.coimage f ≅ image f := is_image.iso_ext (coimage_strong_epi_mono_factorisation f).to_mono_is_image (image.is_image f) +lemma coimage_iso_image'_hom : + (coimage_iso_image' f).hom = cokernel.desc _ (factor_thru_image f) + (by simp [←cancel_mono (limits.image.ι f)]) := +begin + ext, + simp only [←cancel_mono (limits.image.ι f), is_image.iso_ext_hom, cokernel.π_desc, category.assoc, + is_image.lift_ι, coimage_strong_epi_mono_factorisation_to_mono_factorisation_m, + limits.image.fac], +end + +lemma factor_thru_image_comp_coimage_iso_image'_inv : + factor_thru_image f ≫ (coimage_iso_image' f).inv = cokernel.π _ := +by simp only [is_image.iso_ext_inv, image.is_image_lift, image.fac_lift, + coimage_strong_epi_mono_factorisation_to_mono_factorisation_e] + /-- There is a canonical isomorphism between the abelian image and the categorical image of a morphism. -/ abbreviation image_iso_image : abelian.image f ≅ image f := is_image.iso_ext (image_strong_epi_mono_factorisation f).to_mono_is_image (image.is_image f) +lemma image_iso_image_hom_comp_image_ι : + (image_iso_image f).hom ≫ limits.image.ι _ = kernel.ι _ := +by simp only [is_image.iso_ext_hom, is_image.lift_ι, + image_strong_epi_mono_factorisation_to_mono_factorisation_m] + +lemma image_iso_image_inv : + (image_iso_image f).inv = kernel.lift _ (limits.image.ι f) + (by simp [←cancel_epi (factor_thru_image f)]) := +begin + ext, + simp only [is_image.iso_ext_inv, image.is_image_lift, limits.image.fac_lift, + image_strong_epi_mono_factorisation_to_mono_factorisation_e, category.assoc, + kernel.lift_ι, limits.image.fac], +end + end images section cokernel_of_kernel diff --git a/src/category_theory/abelian/opposite.lean b/src/category_theory/abelian/opposite.lean index bf0c940586623..6b8fe7955ef6c 100644 --- a/src/category_theory/abelian/opposite.lean +++ b/src/category_theory/abelian/opposite.lean @@ -121,6 +121,46 @@ by simp def cokernel_unop_unop : cokernel g.unop ≅ (kernel g).unop := (cokernel_unop_op g).unop.symm +/-- The opposite of the image of `g.unop` is the image of `g.` -/ +def image_unop_op : opposite.op (image g.unop) ≅ image g := +(abelian.image_iso_image _).op ≪≫ (cokernel_op_op _).symm ≪≫ + cokernel_iso_of_eq (cokernel.π_unop _) ≪≫ (cokernel_epi_comp _ _) + ≪≫ (cokernel_comp_is_iso _ _) ≪≫ (abelian.coimage_iso_image' _) + +/-- The opposite of the image of `f` is the image of `f.op`. -/ +def image_op_op : opposite.op (image f) ≅ image f.op := image_unop_op f.op + +/-- The image of `f.op` is the opposite of the image of `f`. -/ +def image_op_unop : (image f.op).unop ≅ image f := (image_unop_op f.op).unop + +/-- The image of `g` is the opposite of the image of `g.unop.` -/ +def image_unop_unop : (image g).unop ≅ image g.unop := (image_unop_op g).unop + +lemma image_ι_op_comp_image_unop_op_hom : + (image.ι g.unop).op ≫ (image_unop_op g).hom = factor_thru_image g := +begin + dunfold image_unop_op, + simp only [←category.assoc, ←op_comp, iso.trans_hom, iso.symm_hom, iso.op_hom, cokernel_op_op_inv, + cokernel_comp_is_iso_hom, cokernel_epi_comp_hom, cokernel_iso_of_eq_hom_comp_desc_assoc, + abelian.coimage_iso_image'_hom, eq_to_hom_refl, is_iso.inv_id, + category.id_comp (cokernel.π (kernel.ι g))], + simp only [category.assoc, abelian.image_iso_image_hom_comp_image_ι, kernel.lift_ι, + quiver.hom.op_unop, cokernel.π_desc], +end + +lemma image_unop_op_hom_comp_image_ι : + (image_unop_op g).hom ≫ image.ι g = (factor_thru_image g.unop).op := +by simp only [←cancel_epi (image.ι g.unop).op, ←category.assoc, image_ι_op_comp_image_unop_op_hom, + ←op_comp, image.fac, quiver.hom.op_unop] + +lemma factor_thru_image_comp_image_unop_op_inv : + factor_thru_image g ≫ (image_unop_op g).inv = (image.ι g.unop).op := +by rw [iso.comp_inv_eq, image_ι_op_comp_image_unop_op_hom] + +lemma image_unop_op_inv_comp_op_factor_thru_image : + (image_unop_op g).inv ≫ (factor_thru_image g.unop).op = image.ι g := +by rw [iso.inv_comp_eq, image_unop_op_hom_comp_image_ι] + end end category_theory From 7a1cc03e365059994e64dc313bb023427bcfb50d Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Thu, 26 Jan 2023 09:45:58 +0000 Subject: [PATCH 0363/1291] =?UTF-8?q?chore(set=5Ftheory/zfc/basic):=20`mem?= =?UTF-8?q?=5Fempty`=20=E2=86=92=20`not=5Fmem=5Fempty`=20(#18262)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Matches `set.not_mem_empty`. --- src/set_theory/zfc/basic.lean | 14 +++++++------- src/set_theory/zfc/ordinal.lean | 2 +- 2 files changed, 8 insertions(+), 8 deletions(-) diff --git a/src/set_theory/zfc/basic.lean b/src/set_theory/zfc/basic.lean index 6482af029f56a..f484ce50eb89b 100644 --- a/src/set_theory/zfc/basic.lean +++ b/src/set_theory/zfc/basic.lean @@ -252,7 +252,7 @@ instance : inhabited pSet := ⟨∅⟩ instance : is_empty (type (∅)) := pempty.is_empty -@[simp] theorem mem_empty (x : pSet.{u}) : x ∉ (∅ : pSet.{u}) := is_empty.exists_iff.1 +@[simp] theorem not_mem_empty (x : pSet.{u}) : x ∉ (∅ : pSet.{u}) := is_empty.exists_iff.1 @[simp] theorem to_set_empty : to_set ∅ = ∅ := by simp [to_set] @@ -528,8 +528,8 @@ protected def empty : Set := mk ∅ instance : has_emptyc Set := ⟨Set.empty⟩ instance : inhabited Set := ⟨∅⟩ -@[simp] theorem mem_empty (x) : x ∉ (∅ : Set.{u}) := -quotient.induction_on x pSet.mem_empty +@[simp] theorem not_mem_empty (x) : x ∉ (∅ : Set.{u}) := +quotient.induction_on x pSet.not_mem_empty @[simp] theorem to_set_empty : to_set ∅ = ∅ := by simp [to_set] @@ -547,8 +547,8 @@ begin end theorem eq_empty (x : Set.{u}) : x = ∅ ↔ ∀ y : Set.{u}, y ∉ x := -⟨λ h y, (h.symm ▸ mem_empty y), -λ h, ext (λ y, ⟨λ yx, absurd yx (h y), λ y0, absurd y0 (mem_empty _)⟩)⟩ +⟨λ h y, (h.symm ▸ not_mem_empty y), +λ h, ext (λ y, ⟨λ yx, absurd yx (h y), λ y0, absurd y0 (not_mem_empty _)⟩)⟩ /-- `insert x y` is the set `{x} ∪ y` -/ protected def insert : Set → Set → Set := @@ -586,7 +586,7 @@ theorem mem_insert_of_mem {y z : Set} (x) (h : z ∈ y): z ∈ insert x y := mem by { ext, simp } @[simp] theorem mem_singleton {x y : Set.{u}} : x ∈ @singleton Set.{u} Set.{u} _ y ↔ x = y := -iff.trans mem_insert_iff ⟨λ o, or.rec (λ h, h) (λ n, absurd n (mem_empty _)) o, or.inl⟩ +iff.trans mem_insert_iff ⟨λ o, or.rec (λ h, h) (λ n, absurd n (not_mem_empty _)) o, or.inl⟩ @[simp] theorem to_set_singleton (x : Set) : ({x} : Set).to_set = {x} := by { ext, simp } @@ -936,7 +936,7 @@ to_Set_of_Set _ _ set.ext $ λ y, Set.mem_sep @[simp] theorem empty_hom : ↑(∅ : Set.{u}) = (∅ : Class.{u}) := -set.ext $ λ y, (iff_false _).2 (Set.mem_empty y) +set.ext $ λ y, (iff_false _).2 (Set.not_mem_empty y) @[simp] theorem insert_hom (x y : Set.{u}) : (@insert Set.{u} Class.{u} _ x y) = ↑(insert x y) := set.ext $ λ z, iff.symm Set.mem_insert_iff diff --git a/src/set_theory/zfc/ordinal.lean b/src/set_theory/zfc/ordinal.lean index b24835b989ce9..c752b5c314a48 100644 --- a/src/set_theory/zfc/ordinal.lean +++ b/src/set_theory/zfc/ordinal.lean @@ -35,7 +35,7 @@ namespace Set /-- A transitive set is one where every element is a subset. -/ def is_transitive (x : Set) : Prop := ∀ y ∈ x, y ⊆ x -@[simp] theorem empty_is_transitive : is_transitive ∅ := λ y hy, (mem_empty y hy).elim +@[simp] theorem empty_is_transitive : is_transitive ∅ := λ y hy, (not_mem_empty y hy).elim theorem is_transitive.subset_of_mem (h : x.is_transitive) : y ∈ x → y ⊆ x := h y From 2ef3da38f8ad9f5e3ae79e6dd2988986d121f6b2 Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Thu, 26 Jan 2023 09:45:59 +0000 Subject: [PATCH 0364/1291] feat(special_functions/integrals): integral of `cos(a * x)` for `a` complex (#18279) --- src/analysis/special_functions/integrals.lean | 16 ++++++++++++++++ 1 file changed, 16 insertions(+) diff --git a/src/analysis/special_functions/integrals.lean b/src/analysis/special_functions/integrals.lean index cedeb1c491154..f9f502040cfb7 100644 --- a/src/analysis/special_functions/integrals.lean +++ b/src/analysis/special_functions/integrals.lean @@ -380,6 +380,22 @@ by rw integral_deriv_eq_sub' (λ x, -cos x); norm_num [continuous_on_sin] lemma integral_cos : ∫ x in a..b, cos x = sin b - sin a := by rw integral_deriv_eq_sub'; norm_num [continuous_on_cos] +lemma integral_cos_mul_complex {z : ℂ} (hz : z ≠ 0) (a b : ℝ) : + ∫ x in a..b, complex.cos (z * x) = complex.sin (z * b) / z - complex.sin (z * a) / z := +begin + apply integral_eq_sub_of_has_deriv_at, + swap, + { apply continuous.interval_integrable, + exact complex.continuous_cos.comp (continuous_const.mul complex.continuous_of_real) }, + intros x hx, + have a := complex.has_deriv_at_sin (↑x * z), + have b : has_deriv_at (λ y, y * z : ℂ → ℂ) z ↑x := has_deriv_at_mul_const _, + have c : has_deriv_at (λ (y : ℂ), complex.sin (y * z)) _ ↑x := has_deriv_at.comp x a b, + convert has_deriv_at.comp_of_real (c.div_const z), + { simp_rw mul_comm }, + { rw [mul_div_cancel _ hz, mul_comm] }, +end + lemma integral_cos_sq_sub_sin_sq : ∫ x in a..b, cos x ^ 2 - sin x ^ 2 = sin b * cos b - sin a * cos a := by simpa only [sq, sub_eq_add_neg, neg_mul_eq_mul_neg] using integral_deriv_mul_eq_sub From 9b94375a7e87e8284db3fff2ef5c02a919b56454 Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Thu, 26 Jan 2023 09:46:00 +0000 Subject: [PATCH 0365/1291] feat(data/real/pi): refactor Wallis formula for pi (#18280) This is preparation for adding Euler's sine product (whose proof is a generalisation of the proof of Wallis' formula) --- src/analysis/special_functions/stirling.lean | 53 ++-- src/data/real/pi/wallis.lean | 279 +++++++++++++++---- 2 files changed, 240 insertions(+), 92 deletions(-) diff --git a/src/analysis/special_functions/stirling.lean b/src/analysis/special_functions/stirling.lean index 01a297e311410..1e4fcd66f44b4 100644 --- a/src/analysis/special_functions/stirling.lean +++ b/src/analysis/special_functions/stirling.lean @@ -20,16 +20,15 @@ The proof follows: . We proceed in two parts. -### Part 1 -We consider the fraction sequence $a_n$ of fractions $\frac{n!}{\sqrt{2n}(\frac{n}{e})^n}$ and -prove that this sequence converges against a real, positive number $a$. For this the two main +**Part 1**: We consider the sequence $a_n$ of fractions $\frac{n!}{\sqrt{2n}(\frac{n}{e})^n}$ +and prove that this sequence converges to a real, positive number $a$. For this the two main ingredients are - taking the logarithm of the sequence and - - use the series expansion of $\log(1 + x)$. + - using the series expansion of $\log(1 + x)$. -### Part 2 -We use the fact that the series defined in part 1 converges againt a real number $a$ and prove that -$a = \sqrt{\pi}$. Here the main ingredient is the convergence of the Wallis product. +**Part 2**: We use the fact that the series defined in part 1 converges againt a real number $a$ +and prove that $a = \sqrt{\pi}$. Here the main ingredient is the convergence of Wallis' product +formula for `π`. -/ open_locale topological_space real big_operators nat @@ -202,26 +201,6 @@ end https://proofwiki.org/wiki/Stirling%27s_Formula#Part_2 -/ -/-- For `n : ℕ`, define `w n` as `2^(4*n) * n!^4 / ((2*n)!^2 * (2*n + 1))` -/ -noncomputable def w (n : ℕ) : ℝ := -(2 ^ (4 * n) * n! ^ 4) / ((2 * n)!^ 2 * (2 * n + 1)) - -/-- The sequence `w n` converges to `π/2` -/ -lemma tendsto_w_at_top: tendsto (λ (n : ℕ), w n) at_top (𝓝 (π/2)) := -begin - convert tendsto_prod_pi_div_two, - funext n, - induction n with n ih, - { rw [w, prod_range_zero, cast_zero, mul_zero, pow_zero, one_mul, mul_zero, factorial_zero, - cast_one, one_pow, one_pow, one_mul, mul_zero, zero_add, div_one] }, - rw [w, prod_range_succ, ←ih, w, _root_.div_mul_div_comm, _root_.div_mul_div_comm], - refine (div_eq_div_iff _ _).mpr _, - any_goals { exact ne_of_gt (by positivity) }, - simp_rw [nat.mul_succ, factorial_succ, pow_succ], - push_cast, - ring_nf, -end - /-- The sequence `n / (2 * n + 1)` tends to `1/2` -/ lemma tendsto_self_div_two_mul_self_add_one : tendsto (λ (n : ℕ), (n : ℝ) / (2 * n + 1)) at_top (𝓝 (1 / 2)) := @@ -229,16 +208,16 @@ begin conv { congr, skip, skip, rw [one_div, ←add_zero (2 : ℝ)] }, refine (((tendsto_const_div_at_top_nhds_0_nat 1).const_add (2 : ℝ)).inv₀ ((add_zero (2 : ℝ)).symm ▸ two_ne_zero)).congr' (eventually_at_top.mpr ⟨1, λ n hn, _⟩), - rw [add_div' (1 : ℝ) (2 : ℝ) (n : ℝ) (cast_ne_zero.mpr (one_le_iff_ne_zero.mp hn)), inv_div], + rw [add_div' (1 : ℝ) 2 n (cast_ne_zero.mpr (one_le_iff_ne_zero.mp hn)), inv_div], end - /-- For any `n ≠ 0`, we have the identity -`(stirling_seq n)^4/(stirling_seq (2*n))^2 * (n / (2 * n + 1)) = w n`. -/ +`(stirling_seq n)^4 / (stirling_seq (2*n))^2 * (n / (2 * n + 1)) = W n`, where `W n` is the +`n`-th partial product of Wallis' formula for `π / 2`. -/ lemma stirling_seq_pow_four_div_stirling_seq_pow_two_eq (n : ℕ) (hn : n ≠ 0) : - ((stirling_seq n) ^ 4 / (stirling_seq (2 * n)) ^ 2) * (n / (2 * n + 1)) = w n := + ((stirling_seq n) ^ 4 / (stirling_seq (2 * n)) ^ 2) * (n / (2 * n + 1)) = wallis.W n := begin - rw [bit0_eq_two_mul, stirling_seq, pow_mul, stirling_seq, w], + rw [bit0_eq_two_mul, stirling_seq, pow_mul, stirling_seq, wallis.W_eq_factorial_ratio], simp_rw [div_pow, mul_pow], rw [sq_sqrt, sq_sqrt], any_goals { positivity }, @@ -253,10 +232,10 @@ end /-- Suppose the sequence `stirling_seq` (defined above) has the limit `a ≠ 0`. -Then the sequence `w` has limit `a^2/2` +Then the Wallis sequence `W n` has limit `a^2 / 2`. -/ lemma second_wallis_limit (a : ℝ) (hane : a ≠ 0) (ha : tendsto stirling_seq at_top (𝓝 a)) : - tendsto w at_top (𝓝 (a ^ 2 / 2)):= + tendsto wallis.W at_top (𝓝 (a ^ 2 / 2)):= begin refine tendsto.congr' (eventually_at_top.mpr ⟨1, λ n hn, stirling_seq_pow_four_div_stirling_seq_pow_two_eq n (one_le_iff_ne_zero.mp hn)⟩) _, @@ -272,9 +251,9 @@ end theorem tendsto_stirling_seq_sqrt_pi : tendsto (λ (n : ℕ), stirling_seq n) at_top (𝓝 (sqrt π)) := begin obtain ⟨a, hapos, halimit⟩ := stirling_seq_has_pos_limit_a, - have hπ : π / 2 = a ^ 2 / 2 := tendsto_nhds_unique tendsto_w_at_top - (second_wallis_limit a (ne_of_gt hapos) halimit), - rwa [(div_left_inj' (show (2 : ℝ) ≠ 0, from two_ne_zero)).mp hπ, sqrt_sq hapos.le], + have hπ : π / 2 = a ^ 2 / 2 := tendsto_nhds_unique wallis.tendsto_W_nhds_pi_div_two + (second_wallis_limit a hapos.ne' halimit), + rwa [(div_left_inj' (two_ne_zero' ℝ)).mp hπ, sqrt_sq hapos.le], end end stirling diff --git a/src/data/real/pi/wallis.lean b/src/data/real/pi/wallis.lean index 0795c52c2b56a..7eb62616e6fb5 100644 --- a/src/data/real/pi/wallis.lean +++ b/src/data/real/pi/wallis.lean @@ -5,70 +5,239 @@ Authors: Hanting Zhang -/ import analysis.special_functions.integrals -/-! ### The Wallis Product for Pi -/ +/-! # The Wallis formula for Pi -namespace real +This file establishes the Wallis product for `π` (`real.tendsto_prod_pi_div_two`). Our proof is +largely about analyzing the behaviour of the sequence `∫ x in 0..π, sin x ^ n` as `n → ∞`. +See: https://en.wikipedia.org/wiki/Wallis_product + +The proof can be broken down into two pieces. The first step (carried out in +`analysis.special_functions.integrals`) is to use repeated integration by parts to obtain an +explicit formula for this integral, which is rational if `n` is odd and a rational multiple of `π` +if `n` is even. + +The second step, carried out here, is to estimate the ratio +`∫ (x : ℝ) in 0..π, sin x ^ (2 * k + 1) / ∫ (x : ℝ) in 0..π, sin x ^ (2 * k)` and prove that +it converges to one using the squeeze theorem. The final product for `π` is obtained after some +algebraic manipulation. + +## Main statements -open_locale real topological_space big_operators +* `real.wallis.W`: the product of the first `k` terms in Wallis' formula for `π`. +* `real.wallis.W_eq_integral_sin_pow_div_integral_sin_pow`: express `W n` as a ratio of integrals. +* `real.wallis.W_le` and `real.wallis.le_W`: upper and lower bounds for `W n`. +* `real.wallis.integral_sin_pow_odd_sq_eq` and `real.wallis.integral_sin_pow_even_sq_eq`: formulas + for `(∫ x in 0..π, sin x ^ n) ^ 2` in terms of `W`. +* `integral_sin_pow_le` and `le_integral_sin_pow`: bounds for `∫ x in 0..π, sin x ^ n`. +* `real.tendsto_prod_pi_div_two`: the Wallis product formula. + -/ + +open_locale real topological_space big_operators nat open filter finset interval_integral -lemma integral_sin_pow_div_tendsto_one : - tendsto (λ k, (∫ x in 0..π, sin x ^ (2 * k + 1)) / ∫ x in 0..π, sin x ^ (2 * k)) at_top (𝓝 1) := +namespace real + +namespace wallis + +/-- The product of the first `k` terms in Wallis' formula for `π`. -/ +noncomputable def W (k : ℕ) : ℝ := +∏ i in range k, (2 * i + 2) / (2 * i + 1) * ((2 * i + 2) / (2 * i + 3)) + +lemma W_succ (k : ℕ) : + W (k + 1) = W k * ((2 * k + 2) / (2 * k + 1) * ((2 * k + 2) / (2 * k + 3))) := +prod_range_succ _ _ + +lemma W_pos (k : ℕ) : 0 < W k := +begin + induction k with k hk, + { unfold W, simp }, + { rw W_succ, + refine mul_pos hk (mul_pos (div_pos _ _) (div_pos _ _)); + positivity } +end + +lemma W_eq_factorial_ratio (n : ℕ) : + W n = (2 ^ (4 * n) * n! ^ 4) / ((2 * n)!^ 2 * (2 * n + 1)) := begin - have h₃ : ∀ n, (∫ x in 0..π, sin x ^ (2 * n + 1)) / ∫ x in 0..π, sin x ^ (2 * n) ≤ 1 := - λ n, (div_le_one (integral_sin_pow_pos _)).mpr (integral_sin_pow_succ_le _), - have h₄ : - ∀ n, (∫ x in 0..π, sin x ^ (2 * n + 1)) / ∫ x in 0..π, sin x ^ (2 * n) ≥ 2 * n / (2 * n + 1), - { rintro ⟨n⟩, - { have : 0 ≤ (1 + 1) / π, exact div_nonneg (by norm_num) pi_pos.le, - simp [this] }, - calc (∫ x in 0..π, sin x ^ (2 * n.succ + 1)) / ∫ x in 0..π, sin x ^ (2 * n.succ) ≥ - (∫ x in 0..π, sin x ^ (2 * n.succ + 1)) / ∫ x in 0..π, sin x ^ (2 * n + 1) : - by { refine div_le_div (integral_sin_pow_pos _).le le_rfl (integral_sin_pow_pos _) _, - convert integral_sin_pow_succ_le (2 * n + 1) using 1 } - ... = 2 * ↑(n.succ) / (2 * ↑(n.succ) + 1) : - by { rw div_eq_iff (integral_sin_pow_pos (2 * n + 1)).ne', - convert integral_sin_pow (2 * n + 1), simp with field_simps, norm_cast } }, - refine tendsto_of_tendsto_of_tendsto_of_le_of_le _ _ (λ n, (h₄ n).le) (λ n, (h₃ n)), - { refine metric.tendsto_at_top.mpr (λ ε hε, ⟨⌈1 / ε⌉₊, λ n hn, _⟩), - have h : (2:ℝ) * n / (2 * n + 1) - 1 = -1 / (2 * n + 1), - { conv_lhs { congr, skip, rw ← @div_self _ _ ((2:ℝ) * n + 1) (by { norm_cast, linarith }), }, - rw [← sub_div, ← sub_sub, sub_self, zero_sub] }, - have hpos : (0:ℝ) < 2 * n + 1, { norm_cast, norm_num }, - rw [dist_eq, h, abs_div, abs_neg, abs_one, abs_of_pos hpos, one_div_lt hpos hε], - calc 1 / ε ≤ ⌈1 / ε⌉₊ : nat.le_ceil _ - ... ≤ n : by exact_mod_cast hn.le - ... < 2 * n + 1 : by { norm_cast, linarith } }, - { exact tendsto_const_nhds }, + induction n with n IH, + { simp only [W, prod_range_zero, nat.factorial_zero, mul_zero, pow_zero, algebra_map.coe_one, + one_pow, mul_one, algebra_map.coe_zero, zero_add, div_self, ne.def, one_ne_zero, + not_false_iff] }, + { unfold W at ⊢ IH, + rw [prod_range_succ, IH, _root_.div_mul_div_comm, _root_.div_mul_div_comm], + refine (div_eq_div_iff _ _).mpr _, + any_goals { exact ne_of_gt (by positivity) }, + simp_rw [nat.mul_succ, nat.factorial_succ, pow_succ], + push_cast, + ring_nf } end -/-- This theorem establishes the Wallis Product for `π`. Our proof is largely about analyzing - the behavior of the ratio of the integral of `sin x ^ n` as `n → ∞`. - See: https://en.wikipedia.org/wiki/Wallis_product - - The proof can be broken down into two pieces. - (Pieces involving general properties of the integral of `sin x ^n` can be found - in `analysis.special_functions.integrals`.) First, we use integration by parts to obtain a - recursive formula for `∫ x in 0..π, sin x ^ (n + 2)` in terms of `∫ x in 0..π, sin x ^ n`. - From this we can obtain closed form products of `∫ x in 0..π, sin x ^ (2 * n)` and - `∫ x in 0..π, sin x ^ (2 * n + 1)` via induction. Next, we study the behavior of the ratio - `∫ (x : ℝ) in 0..π, sin x ^ (2 * k + 1)) / ∫ (x : ℝ) in 0..π, sin x ^ (2 * k)` and prove that - it converges to one using the squeeze theorem. The final product for `π` is obtained after some - algebraic manipulation. -/ -theorem tendsto_prod_pi_div_two : - tendsto (λ k, ∏ i in range k, - (((2:ℝ) * i + 2) / (2 * i + 1)) * ((2 * i + 2) / (2 * i + 3))) at_top (𝓝 (π/2)) := +lemma W_eq_integral_sin_pow_div_integral_sin_pow (k : ℕ) : + (π/2)⁻¹ * W k = (∫ (x : ℝ) in 0..π, sin x ^ (2 * k + 1)) / ∫ (x : ℝ) in 0..π, sin x ^ (2 * k) := begin - suffices h : tendsto (λ k, (π / 2)⁻¹ * ∏ i in range k, - (2 * i + 2) / (2 * i + 1) * ((2 * i + 2) / (2 * i + 3))) at_top (𝓝 1), - { convert h.const_mul (π / 2), - { simp_rw mul_inv_cancel_left₀ (show π / 2 ≠ 0, by norm_num [pi_ne_zero]) }, - { rw mul_one } }, - convert integral_sin_pow_div_tendsto_one, - funext, rw [integral_sin_pow_even, integral_sin_pow_odd, mul_div_mul_comm, ←prod_div_distrib, inv_div], - congr' with i, - rw [div_div_div_comm, div_div_eq_mul_div, mul_div_assoc], + simp_rw [div_div_div_comm, div_div_eq_mul_div, mul_div_assoc], + refl, +end + +lemma W_le (k : ℕ) : W k ≤ π / 2 := +begin + rw [←div_le_one pi_div_two_pos, div_eq_inv_mul], + rw [W_eq_integral_sin_pow_div_integral_sin_pow, div_le_one (integral_sin_pow_pos _)], + apply integral_sin_pow_succ_le, +end + +lemma le_W (k : ℕ) : ((2:ℝ) * k + 1) / (2 * k + 2) * (π / 2) ≤ W k := +begin + rw [←le_div_iff pi_div_two_pos, div_eq_inv_mul (W k) _], + rw [W_eq_integral_sin_pow_div_integral_sin_pow, le_div_iff (integral_sin_pow_pos _)], + convert integral_sin_pow_succ_le (2 * k + 1), + rw [integral_sin_pow (2 * k)], + simp only [sin_zero, zero_pow', ne.def, nat.succ_ne_zero, not_false_iff, zero_mul, sin_pi, + tsub_zero, nat.cast_mul, nat.cast_bit0, algebra_map.coe_one, zero_div, zero_add], +end + +lemma tendsto_W_nhds_pi_div_two : tendsto W at_top (𝓝 $ π / 2) := +begin + refine tendsto_of_tendsto_of_tendsto_of_le_of_le _ tendsto_const_nhds le_W W_le, + have : 𝓝 (π / 2) = 𝓝 ((1 - 0) * (π / 2)), by rw [sub_zero, one_mul], rw this, + refine tendsto.mul _ tendsto_const_nhds, + have h : ∀ (n:ℕ), ((2:ℝ) * n + 1) / (2 * n + 2) = 1 - 1 / (2 * n + 2), + { intro n, + rw [sub_div' _ _ _ (ne_of_gt (add_pos_of_nonneg_of_pos + (mul_nonneg ((two_pos : 0 < (2:ℝ)).le) (nat.cast_nonneg _)) two_pos)), one_mul], + congr' 1, ring }, + simp_rw h, + refine (tendsto_const_nhds.div_at_top _).const_sub _, + refine tendsto.at_top_add _ tendsto_const_nhds, + exact tendsto_coe_nat_at_top_at_top.const_mul_at_top two_pos, +end + +lemma W_eq_mul_sq (k : ℕ) : + W k = (2 * k + 1) * (∏ i in range k, ((2:ℝ) * i + 2) / (2 * i + 3)) ^ 2 := +begin + induction k with k hk, + { simp [W], }, + { unfold W at *, + rw [prod_range_succ, prod_range_succ, hk], + suffices : ∀ (x : ℝ), + (2 * ↑k + 1) * x ^ 2 * ((2 * ↑k + 2) / (2 * ↑k + 1) * ((2 * ↑k + 2) / (2 * ↑k + 3))) = + (2 * ↑(k.succ) + 1) * (x * ((2 * ↑k + 2) / (2 * ↑k + 3))) ^ 2, + { rw this }, + intro x, + have a : (2 * ↑k + 1 : ℝ) ≠ 0, by positivity, + have b : (2 * ↑k + 3 : ℝ) ≠ 0, by positivity, + field_simp, ring } end +lemma integral_sin_pow_odd_sq_eq (k : ℕ) : + (∫ x in 0..π, sin x ^ (2 * k + 1)) ^ 2 = 4 * W k / (2 * k + 1) := +begin + rw integral_sin_pow_odd, + have B := W_eq_mul_sq k, + rw [mul_comm (2 * (k:ℝ) + 1) _, ←div_eq_iff] at B, + { rw [mul_pow, ←B], + ring }, + { positivity }, +end + +lemma integral_sin_pow_even_sq_eq (k : ℕ) : + (∫ x in 0..π, sin x ^ (2 * k)) ^ 2 = π ^ 2 / (2 * k + 1) / W k := +begin + induction k with k hk, + { dsimp only [W], + simp }, + { have np : 0 < 2 * (k:ℝ) + 1, by positivity, + rw [nat.succ_eq_add_one, mul_add 2 k 1, mul_one, integral_sin_pow, sin_zero, sin_pi, + zero_pow (nat.add_pos_right _ zero_lt_one), zero_mul, zero_mul, sub_zero, zero_div, zero_add, + mul_pow, hk, W_succ, nat.cast_add_one, nat.cast_mul, mul_add, mul_one, + add_assoc (2 * (k:ℝ)) 2 1, (by ring : (2:ℝ) + 1 = 3), sq], + have np2 : 2 * (k:ℝ) + 2 ≠ 0, by positivity, + have np3 : 2 * (k:ℝ) + 3 ≠ 0, by positivity, + field_simp [np.ne', (W_pos k).ne'], + ring } +end + +end wallis + end real + +open real real.wallis + +section integral_sin_pow_bounds +/-! ## Bounds for integrals of `sin x ^ n` + +Explicit `O(1/√n)` bounds for `∫ x in 0..π, sin x ^ n`, as a by-product of the proof of Wallis' +formula for `π`. -/ + +lemma integral_sin_pow_odd_le (n : ℕ) : + ∫ x in 0..π, sin x ^ (2 * n + 1) ≤ sqrt (2 * π / (2 * n + 1)) := +begin + have np : 0 < 2 * (n:ℝ) + 1, by positivity, + rw [le_sqrt (integral_sin_pow_pos _).le (div_pos two_pi_pos np).le, integral_sin_pow_odd_sq_eq], + apply div_le_div_of_le np.le, + rw ←le_div_iff' (by norm_num : 0 < (4:ℝ)), + convert W_le n using 1, + ring, +end + +lemma integral_sin_pow_even_le (n : ℕ) : + ∫ x in 0..π, sin x ^ (2 * n) ≤ sqrt (2 * π * (2 * n + 2) / (2 * n + 1) ^ 2) := +begin + have np : 0 < 2 * (n:ℝ) + 1, by positivity, + have np' : 0 < 2 * (n:ℝ) + 2, by positivity, + rw le_sqrt (integral_sin_pow_pos _).le, + swap, { refine div_nonneg _ (sq_nonneg _), exact mul_nonneg (two_pi_pos).le np'.le }, + rw [integral_sin_pow_even_sq_eq, div_le_iff (W_pos n), ←div_le_iff'], + swap, { refine div_pos _ (sq_pos_of_pos np), exact mul_pos two_pi_pos np' }, + convert le_W n, + field_simp [np.ne', np'.ne', pi_pos.ne'], + ring, +end + +lemma integral_sin_pow_le {n : ℕ} (hn : n ≠ 0) : ∫ x in 0..π, sin x ^ n ≤ sqrt (2 * π / n) := +begin + -- this is a slightly weaker bound than `integral_sin_pow_even_le` for even `n`, but uniform in + -- its statement + obtain ⟨k, hk⟩ := nat.even_or_odd' n, + rcases hk with rfl | rfl, + { refine le_trans (integral_sin_pow_even_le k) _, + apply sqrt_le_sqrt, + rw [div_le_div_iff, mul_assoc, mul_le_mul_left two_pi_pos], + rotate, { positivity }, { positivity }, + have : (2 * (k:ℝ) + 2) * ((2 * k : ℕ) : ℝ) = (2 * k + 1) ^ 2 - 1, + { push_cast, ring, }, + rw [this, sub_le_self_iff], + exact zero_le_one }, + { convert integral_sin_pow_odd_le k using 3, + rw [nat.cast_add, nat.cast_mul, nat.cast_two, nat.cast_one] }, +end + +lemma le_integral_sin_pow (n : ℕ) : sqrt (2 * π / (n + 1)) ≤ ∫ x in 0..π, sin x ^ n := +begin + refine sqrt_le_iff.mpr ⟨(integral_sin_pow_pos _).le, _⟩, + obtain ⟨k, hk⟩ := nat.even_or_odd' n, + have np : 0 < 2 * (k:ℝ) + 1, by positivity, + have np' : 2 * (k:ℝ) + 2 ≠ 0, by positivity, + rcases hk with rfl | rfl, + { rw [integral_sin_pow_even_sq_eq, le_div_iff (W_pos _), nat.cast_mul, nat.cast_two, + ←le_div_iff' (div_pos two_pi_pos np)], + convert W_le k using 1, + field_simp [np.ne', np', pi_pos.ne'], + ring }, + { rw [nat.cast_add, nat.cast_mul, nat.cast_two, nat.cast_one, + (by ring : (2:ℝ) * k + 1 + 1 = 2 * k + 2), integral_sin_pow_odd_sq_eq, le_div_iff np, + ←div_le_iff' (by positivity : 0 < (4:ℝ))], + convert le_W k, + field_simp [np.ne', np'], + ring }, +end + +end integral_sin_pow_bounds + +/-- Wallis' product formula for `π / 2`. -/ +theorem real.tendsto_prod_pi_div_two : + tendsto + (λ k, ∏ i in range k, (((2:ℝ) * i + 2) / (2 * i + 1)) * ((2 * i + 2) / (2 * i + 3))) + at_top (𝓝 (π/2)) := +tendsto_W_nhds_pi_div_two From cd45082b33bd2d8dfab54576a31a9af26cccd08d Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Thu, 26 Jan 2023 09:46:02 +0000 Subject: [PATCH 0366/1291] feat(set_theory/zfc/basic): `pSet.equiv.comm` (#18297) --- src/set_theory/zfc/basic.lean | 3 +++ 1 file changed, 3 insertions(+) diff --git a/src/set_theory/zfc/basic.lean b/src/set_theory/zfc/basic.lean index f484ce50eb89b..bdb8336eaa245 100644 --- a/src/set_theory/zfc/basic.lean +++ b/src/set_theory/zfc/basic.lean @@ -132,6 +132,9 @@ pSet.rec_on x $ λ α A IH y, pSet.cases_on y $ λ β B ⟨γ, Γ⟩ ⟨αβ, β @[symm] protected theorem equiv.symm {x y} : equiv x y → equiv y x := (equiv.refl y).euc +protected theorem equiv.comm {x y} : equiv x y ↔ equiv y x := +⟨equiv.symm, equiv.symm⟩ + @[trans] protected theorem equiv.trans {x y z} (h1 : equiv x y) (h2 : equiv y z) : equiv x z := h1.euc h2.symm From f7fc89d5d5ff1db2d1242c7bb0e9062ce47ef47c Mon Sep 17 00:00:00 2001 From: Paul Lezeau <72892199+Paul-Lez@users.noreply.github.com> Date: Thu, 26 Jan 2023 11:48:34 +0000 Subject: [PATCH 0367/1291] feat(number_theory/kummer_dedekind): the return of the Dedekind-Kummer theorem. (#16870) In this PR, we finish the proof of the general case of the Dedekind-Kummer theorem as stated in Neukirch, completing the work done in #15000. Co-authored-by: Paul Lezeau --- .../minpoly/is_integrally_closed.lean | 11 ++ src/number_theory/kummer_dedekind.lean | 111 +++++++++++------- src/ring_theory/adjoin_root.lean | 31 +++++ src/ring_theory/is_adjoin_root.lean | 36 ++++++ 4 files changed, 145 insertions(+), 44 deletions(-) diff --git a/src/field_theory/minpoly/is_integrally_closed.lean b/src/field_theory/minpoly/is_integrally_closed.lean index 37c45db1fdf8e..4f6466607268e 100644 --- a/src/field_theory/minpoly/is_integrally_closed.lean +++ b/src/field_theory/minpoly/is_integrally_closed.lean @@ -140,6 +140,17 @@ begin { rw [(monic hs).leading_coeff, hmo.leading_coeff] } end +theorem prime_of_is_integrally_closed {x : S} (hx : is_integral R x) : + _root_.prime (minpoly R x) := +begin + refine ⟨(minpoly.monic hx).ne_zero, ⟨by by_contra h_contra ; + exact (ne_of_lt (minpoly.degree_pos hx)) (degree_eq_zero_of_is_unit h_contra).symm, + λ a b h, or_iff_not_imp_left.mpr (λ h', _)⟩⟩, + rw ← minpoly.is_integrally_closed_dvd_iff hx at ⊢ h' h, + rw aeval_mul at h, + exact eq_zero_of_ne_zero_of_mul_left_eq_zero h' h, +end + section adjoin_root noncomputable theory diff --git a/src/number_theory/kummer_dedekind.lean b/src/number_theory/kummer_dedekind.lean index 04577ebe05292..b15acf6f64974 100644 --- a/src/number_theory/kummer_dedekind.lean +++ b/src/number_theory/kummer_dedekind.lean @@ -4,9 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Anne Baanen, Paul Lezeau -/ -import ring_theory.adjoin_root -import ring_theory.dedekind_domain.ideal import ring_theory.algebra_tower +import ring_theory.dedekind_domain.ideal +import ring_theory.is_adjoin_root /-! # Kummer-Dedekind theorem @@ -78,6 +78,12 @@ lemma conductor_subset_adjoin : (conductor R x : set S) ⊆ R := lemma mem_conductor_iff {y : S} : y ∈ conductor R x ↔ ∀ (b : S), y * b ∈ R := ⟨λ h, h, λ h, h⟩ +lemma conductor_eq_top_of_adjoin_eq_top (h : R = ⊤) : conductor R x = ⊤ := +by simp only [ideal.eq_top_iff_one, mem_conductor_iff, h, mem_top, forall_const] + +lemma conductor_eq_top_of_power_basis (pb : power_basis R S) : conductor R pb.gen = ⊤ := +conductor_eq_top_of_adjoin_eq_top pb.adjoin_gen_eq_top + variables {I : ideal R} /-- This technical lemma tell us that if `C` is the conductor of `R` and `I` is an ideal of `R` @@ -202,8 +208,9 @@ namespace kummer_dedekind open_locale big_operators polynomial classical -variables [is_domain R] [is_domain S] [is_dedekind_domain S] -variables (pb : power_basis R S) +variables [is_domain R] [is_integrally_closed R] +variables [is_domain S] [is_dedekind_domain S] +variable [no_zero_smul_divisors R S] local attribute [instance] ideal.quotient.field @@ -211,27 +218,37 @@ local attribute [instance] ideal.quotient.field factors of `I*S` are in bijection with those of the minimal polynomial of the generator of `S` over `R`, taken `mod I`.-/ noncomputable def normalized_factors_map_equiv_normalized_factors_min_poly_mk (hI : is_maximal I) - (hI' : I ≠ ⊥) : {J : ideal S | J ∈ normalized_factors (I.map (algebra_map R S) )} ≃ - {d : (R ⧸ I)[X] | d ∈ normalized_factors (map I^.quotient.mk (minpoly R pb.gen)) } := -((normalized_factors_equiv_of_quot_equiv ↑(pb.quotient_equiv_quotient_minpoly_map I) + (hI' : I ≠ ⊥) (hx : (conductor R x).comap (algebra_map R S) ⊔ I = ⊤) + (hx' : is_integral R x) : + {J : ideal S | J ∈ normalized_factors (I.map (algebra_map R S) )} ≃ + {d : (R ⧸ I)[X] | d ∈ normalized_factors (map I^.quotient.mk (minpoly R x))} := +(normalized_factors_equiv_of_quot_equiv + ((quot_adjoin_equiv_quot_map hx + (by { apply no_zero_smul_divisors.algebra_map_injective (algebra.adjoin R {x}) S, + exact subalgebra.no_zero_smul_divisors_top (algebra.adjoin R {x}) })).symm.trans + (((algebra.adjoin.power_basis' hx').quotient_equiv_quotient_minpoly_map I).to_ring_equiv.trans + (quot_equiv_of_eq (show (ideal.span ({(minpoly R (algebra.adjoin.power_basis' hx').gen).map + I^.quotient.mk})) = (ideal.span ({(minpoly R x).map I^.quotient.mk})), + by rw algebra.adjoin.power_basis'_minpoly_gen hx')))) --show that `I * S` ≠ ⊥ (show I.map (algebra_map R S) ≠ ⊥, - by rwa [ne.def, map_eq_bot_iff_of_injective pb.basis.algebra_map_injective, ← ne.def]) + by rwa [ne.def, map_eq_bot_iff_of_injective (no_zero_smul_divisors.algebra_map_injective R S), + ← ne.def]) --show that the ideal spanned by `(minpoly R pb.gen) mod I` is non-zero - (by {by_contra, exact (show (map I^.quotient.mk (minpoly R pb.gen) ≠ 0), from - polynomial.map_monic_ne_zero (minpoly.monic pb.is_integral_gen)) - (span_singleton_eq_bot.mp h) } )).trans + (by {by_contra, exact (show (map I^.quotient.mk (minpoly R x) ≠ 0), from + polynomial.map_monic_ne_zero (minpoly.monic hx')) (span_singleton_eq_bot.mp h) } )).trans (normalized_factors_equiv_span_normalized_factors - (show (map I^.quotient.mk (minpoly R pb.gen)) ≠ 0, from - polynomial.map_monic_ne_zero (minpoly.monic pb.is_integral_gen))).symm) + (show (map I^.quotient.mk (minpoly R x)) ≠ 0, from + polynomial.map_monic_ne_zero (minpoly.monic hx'))).symm /-- The second half of the **Kummer-Dedekind Theorem** in the monogenic case, stating that the bijection `factors_equiv'` defined in the first half preserves multiplicities. -/ -theorem multiplicity_factors_map_eq_multiplicity (hI : is_maximal I) (hI' : I ≠ ⊥) {J : ideal S} - (hJ : J ∈ normalized_factors (I.map (algebra_map R S))) : +theorem multiplicity_factors_map_eq_multiplicity (hI : is_maximal I) (hI' : I ≠ ⊥) + (hx : (conductor R x).comap (algebra_map R S) ⊔ I = ⊤) (hx' : is_integral R x) + {J : ideal S} (hJ : J ∈ normalized_factors (I.map (algebra_map R S))) : multiplicity J (I.map (algebra_map R S)) = - multiplicity ↑(normalized_factors_map_equiv_normalized_factors_min_poly_mk pb hI hI' ⟨J, hJ⟩) - (map I^.quotient.mk (minpoly R pb.gen)) := + multiplicity ↑(normalized_factors_map_equiv_normalized_factors_min_poly_mk hI hI' hx hx' + ⟨J, hJ⟩) (map I^.quotient.mk (minpoly R x)) := by rw [normalized_factors_map_equiv_normalized_factors_min_poly_mk, equiv.coe_trans, function.comp_app, multiplicity_normalized_factors_equiv_span_normalized_factors_symm_eq_multiplicity, @@ -239,9 +256,13 @@ by rw [normalized_factors_map_equiv_normalized_factors_min_poly_mk, equiv.coe_tr /-- The **Kummer-Dedekind Theorem**. -/ theorem normalized_factors_ideal_map_eq_normalized_factors_min_poly_mk_map (hI : is_maximal I) - (hI' : I ≠ ⊥) : normalized_factors (I.map (algebra_map R S)) = multiset.map - (λ f, ((normalized_factors_map_equiv_normalized_factors_min_poly_mk pb hI hI').symm f : ideal S)) - (normalized_factors (polynomial.map I^.quotient.mk (minpoly R pb.gen))).attach := + (hI' : I ≠ ⊥) (hx : (conductor R x).comap (algebra_map R S) ⊔ I = ⊤) + (hx' : is_integral R x) : + normalized_factors (I.map (algebra_map R S)) = + multiset.map + (λ f, ((normalized_factors_map_equiv_normalized_factors_min_poly_mk hI hI' hx hx').symm f : + ideal S)) + (normalized_factors (polynomial.map I^.quotient.mk (minpoly R x))).attach := begin ext J, -- WLOG, assume J is a normalized factor @@ -250,10 +271,10 @@ begin simp only [multiset.mem_attach, true_and, not_exists], rintros J' rfl, exact hJ - ((normalized_factors_map_equiv_normalized_factors_min_poly_mk pb hI hI').symm J').prop }, + ((normalized_factors_map_equiv_normalized_factors_min_poly_mk hI hI' hx hx').symm J').prop }, -- Then we just have to compare the multiplicities, which we already proved are equal. - have := multiplicity_factors_map_eq_multiplicity pb hI hI' hJ, + have := multiplicity_factors_map_eq_multiplicity hI hI' hx hx' hJ, rw [multiplicity_eq_count_normalized_factors, multiplicity_eq_count_normalized_factors, unique_factorization_monoid.normalize_normalized_factor _ hJ, unique_factorization_monoid.normalize_normalized_factor, @@ -261,45 +282,47 @@ begin at this, refine this.trans _, -- Get rid of the `map` by applying the equiv to both sides. - generalize hJ' : (normalized_factors_map_equiv_normalized_factors_min_poly_mk pb hI hI') + generalize hJ' : (normalized_factors_map_equiv_normalized_factors_min_poly_mk hI hI' hx hx') ⟨J, hJ⟩ = J', - have : ((normalized_factors_map_equiv_normalized_factors_min_poly_mk pb hI hI').symm - J' : ideal S) = J, + have : ((normalized_factors_map_equiv_normalized_factors_min_poly_mk hI hI' hx hx').symm J' : + ideal S) = J, { rw [← hJ', equiv.symm_apply_apply _ _, subtype.coe_mk] }, subst this, -- Get rid of the `attach` by applying the subtype `coe` to both sides. rw [multiset.count_map_eq_count' (λ f, - ((normalized_factors_map_equiv_normalized_factors_min_poly_mk pb hI hI').symm f + ((normalized_factors_map_equiv_normalized_factors_min_poly_mk hI hI' hx hx').symm f : ideal S)), multiset.attach_count_eq_count_coe], { exact subtype.coe_injective.comp (equiv.injective _) }, - { exact (normalized_factors_map_equiv_normalized_factors_min_poly_mk pb hI hI' _).prop }, + { exact (normalized_factors_map_equiv_normalized_factors_min_poly_mk hI hI' hx hx' _).prop}, { exact irreducible_of_normalized_factor _ - (normalized_factors_map_equiv_normalized_factors_min_poly_mk pb hI hI' _).prop }, - { exact polynomial.map_monic_ne_zero (minpoly.monic pb.is_integral_gen) }, + (normalized_factors_map_equiv_normalized_factors_min_poly_mk hI hI' hx hx' _).prop }, + { exact polynomial.map_monic_ne_zero (minpoly.monic hx') }, { exact irreducible_of_normalized_factor _ hJ }, - { rwa [← bot_eq_zero, ne.def, map_eq_bot_iff_of_injective pb.basis.algebra_map_injective] }, + { rwa [← bot_eq_zero, ne.def, map_eq_bot_iff_of_injective + (no_zero_smul_divisors.algebra_map_injective R S)] }, end theorem ideal.irreducible_map_of_irreducible_minpoly (hI : is_maximal I) (hI' : I ≠ ⊥) - (hf : irreducible (map I^.quotient.mk (minpoly R pb.gen))) : + (hx : (conductor R x).comap (algebra_map R S) ⊔ I = ⊤) + (hx' : is_integral R x) (hf : irreducible (map I^.quotient.mk (minpoly R x))) : irreducible (I.map (algebra_map R S)) := begin - have mem_norm_factors : normalize (map I^.quotient.mk (minpoly R pb.gen)) ∈ normalized_factors - (map I^.quotient.mk (minpoly R pb.gen)) := by simp [normalized_factors_irreducible hf], - suffices : ∃ x, normalized_factors (I.map (algebra_map R S)) = {x}, - { obtain ⟨x, hx⟩ := this, - have h := normalized_factors_prod (show I.map (algebra_map R S) ≠ 0, by - rwa [← bot_eq_zero, ne.def, map_eq_bot_iff_of_injective pb.basis.algebra_map_injective]), - rw [associated_iff_eq, hx, multiset.prod_singleton] at h, + have mem_norm_factors : normalize (map I^.quotient.mk (minpoly R x)) ∈ normalized_factors + (map I^.quotient.mk (minpoly R x)) := by simp [normalized_factors_irreducible hf], + suffices : ∃ y, normalized_factors (I.map (algebra_map R S)) = {y}, + { obtain ⟨y, hy⟩ := this, + have h := normalized_factors_prod (show I.map (algebra_map R S) ≠ 0, by rwa [← bot_eq_zero, + ne.def, map_eq_bot_iff_of_injective (no_zero_smul_divisors.algebra_map_injective R S)]), + rw [associated_iff_eq, hy, multiset.prod_singleton] at h, rw ← h, - exact irreducible_of_normalized_factor x - (show x ∈ normalized_factors (I.map (algebra_map R S)), by simp [hx]) }, - rw normalized_factors_ideal_map_eq_normalized_factors_min_poly_mk_map pb hI hI', - use ((normalized_factors_map_equiv_normalized_factors_min_poly_mk pb hI hI').symm - ⟨normalize (map I^.quotient.mk (minpoly R pb.gen)), mem_norm_factors⟩ : ideal S), + exact irreducible_of_normalized_factor y + (show y ∈ normalized_factors (I.map (algebra_map R S)), by simp [hy]) }, + rw normalized_factors_ideal_map_eq_normalized_factors_min_poly_mk_map hI hI' hx hx', + use ((normalized_factors_map_equiv_normalized_factors_min_poly_mk hI hI' hx hx').symm + ⟨normalize (map I^.quotient.mk (minpoly R x)), mem_norm_factors⟩ : ideal S), rw multiset.map_eq_singleton, - use ⟨normalize (map I^.quotient.mk (minpoly R pb.gen)), mem_norm_factors⟩, + use ⟨normalize (map I^.quotient.mk (minpoly R x)), mem_norm_factors⟩, refine ⟨_, rfl⟩, apply multiset.map_injective subtype.coe_injective, rw [multiset.attach_map_coe, multiset.map_singleton, subtype.coe_mk], diff --git a/src/ring_theory/adjoin_root.lean b/src/ring_theory/adjoin_root.lean index 6fb17d1e38b18..445451c46b548 100644 --- a/src/ring_theory/adjoin_root.lean +++ b/src/ring_theory/adjoin_root.lean @@ -168,6 +168,20 @@ by rw [is_root, eval_map, eval₂_root] lemma is_algebraic_root (hf : f ≠ 0) : is_algebraic R (root f) := ⟨f, hf, eval₂_root f⟩ +lemma of.injective_of_degree_ne_zero [is_domain R] (hf : f.degree ≠ 0) : + function.injective (adjoin_root.of f) := +begin + rw injective_iff_map_eq_zero, + intros p hp, + rw [adjoin_root.of, ring_hom.comp_apply, adjoin_root.mk_eq_zero] at hp, + by_cases h : f = 0, + { exact C_eq_zero.mp (eq_zero_of_zero_dvd (by rwa h at hp)) }, + { contrapose! hf with h_contra, + rw ← degree_C h_contra, + apply le_antisymm (degree_le_of_dvd hp (by rwa [ne.def, C_eq_zero])) _, + rwa [degree_C h_contra, zero_le_degree_iff] }, +end + variables [comm_ring S] /-- Lift a ring homomorphism `i : R →+* S` to `adjoin_root f →+* S`. -/ @@ -244,6 +258,23 @@ lemma alg_hom_subsingleton {S : Type*} [comm_ring S] [algebra R S] {r : R} : end adjoin_inv +section prime + +variable {f} + +theorem is_domain_of_prime (hf : prime f) : is_domain (adjoin_root f) := +(ideal.quotient.is_domain_iff_prime (span {f} : ideal R[X])).mpr $ + (ideal.span_singleton_prime hf.ne_zero).mpr hf + +theorem no_zero_smul_divisors_of_prime_of_degree_ne_zero [is_domain R] (hf : prime f) + (hf' : f.degree ≠ 0) : no_zero_smul_divisors R (adjoin_root f) := +begin + haveI := is_domain_of_prime hf, + exact no_zero_smul_divisors.iff_algebra_map_injective.mpr (of.injective_of_degree_ne_zero hf') +end + +end prime + end comm_ring section irreducible diff --git a/src/ring_theory/is_adjoin_root.lean b/src/ring_theory/is_adjoin_root.lean index 766e7a34bfebb..087a07edd462b 100644 --- a/src/ring_theory/is_adjoin_root.lean +++ b/src/ring_theory/is_adjoin_root.lean @@ -290,6 +290,24 @@ protected def is_adjoin_root_monic (hf : monic f) : { monic := hf, .. adjoin_root.is_adjoin_root f } +@[simp] +lemma is_adjoin_root_map_eq_mk : + (adjoin_root.is_adjoin_root f).map = adjoin_root.mk f := rfl + +@[simp] +lemma is_adjoin_root_monic_map_eq_mk (hf : f.monic) : + (adjoin_root.is_adjoin_root_monic f hf).map = adjoin_root.mk f := rfl + +@[simp] +lemma is_adjoin_root_root_eq_root : + (adjoin_root.is_adjoin_root f).root = adjoin_root.root f := +by simp only [is_adjoin_root.root, adjoin_root.root, adjoin_root.is_adjoin_root_map_eq_mk] + +@[simp] +lemma is_adjoin_root_monic_root_eq_root (hf : monic f) : + (adjoin_root.is_adjoin_root_monic f hf).root = adjoin_root.root f := +by simp only [is_adjoin_root.root, adjoin_root.root, adjoin_root.is_adjoin_root_monic_map_eq_mk] + end adjoin_root namespace is_adjoin_root_monic @@ -639,4 +657,22 @@ by convert (associated.mul_left (minpoly R h.root) $ end is_adjoin_root_monic +section algebra + +open adjoin_root is_adjoin_root minpoly power_basis is_adjoin_root_monic algebra + +lemma algebra.adjoin.power_basis'_minpoly_gen [is_domain R] [is_domain S] + [no_zero_smul_divisors R S] [is_integrally_closed R] {x : S} (hx' : is_integral R x) : + minpoly R x = minpoly R (algebra.adjoin.power_basis' hx').gen := +begin + haveI := is_domain_of_prime (prime_of_is_integrally_closed hx'), + haveI := no_zero_smul_divisors_of_prime_of_degree_ne_zero + (prime_of_is_integrally_closed hx') (ne_of_lt (degree_pos hx')).symm, + rw [← minpoly_gen_eq, adjoin.power_basis', minpoly_gen_map, minpoly_gen_eq, power_basis'_gen, + ← is_adjoin_root_monic_root_eq_root _ (monic hx'), minpoly_eq], + exact irreducible hx', +end + +end algebra + end comm_ring From 975c8c329887c50db6f3556a5f382292ee152ff9 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Thu, 26 Jan 2023 21:54:57 +0000 Subject: [PATCH 0368/1291] feat(data/list/lemmas): Range of `foldr` (#13328) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Mathlib4 pair: https://github.com/leanprover-community/mathlib4/pull/1870 Co-authored-by: Eric Wieser Co-authored-by: Yaël Dillies --- src/data/list/lemmas.lean | 36 +++++++++++++++++++++++++++++++++++- 1 file changed, 35 insertions(+), 1 deletion(-) diff --git a/src/data/list/lemmas.lean b/src/data/list/lemmas.lean index afc0d0070e9eb..7705922f783c9 100644 --- a/src/data/list/lemmas.lean +++ b/src/data/list/lemmas.lean @@ -16,7 +16,7 @@ Split out from `data.list.basic` to reduce its dependencies. open list -variables {α β : Type*} +variables {α β γ : Type*} namespace list @@ -66,4 +66,38 @@ begin { simpa [nat.succ_le_succ_iff] using hm } } } end +lemma foldr_range_subset_of_range_subset {f : β → α → α} {g : γ → α → α} + (hfg : set.range f ⊆ set.range g) (a : α) : + set.range (foldr f a) ⊆ set.range (foldr g a) := +begin + rintro _ ⟨l, rfl⟩, + induction l with b l H, + { exact ⟨[], rfl⟩ }, + { cases hfg (set.mem_range_self b) with c hgf, + cases H with m hgf', + rw [foldr_cons, ←hgf, ←hgf'], + exact ⟨c :: m, rfl⟩ } +end + +lemma foldl_range_subset_of_range_subset {f : α → β → α} {g : α → γ → α} + (hfg : set.range (λ a c, f c a) ⊆ set.range (λ b c, g c b)) (a : α) : + set.range (foldl f a) ⊆ set.range (foldl g a) := +begin + change set.range (λ l, _) ⊆ set.range (λ l, _), + simp_rw ←foldr_reverse at hfg ⊢, + simp_rw [set.range_comp _ list.reverse, reverse_involutive.bijective.surjective.range_eq, + set.image_univ], + exact foldr_range_subset_of_range_subset hfg a, +end + +lemma foldr_range_eq_of_range_eq {f : β → α → α} {g : γ → α → α} (hfg : set.range f = set.range g) + (a : α) : + set.range (foldr f a) = set.range (foldr g a) := +(foldr_range_subset_of_range_subset hfg.le a).antisymm (foldr_range_subset_of_range_subset hfg.ge a) + +lemma foldl_range_eq_of_range_eq {f : α → β → α} {g : α → γ → α} + (hfg : set.range (λ a c, f c a) = set.range (λ b c, g c b)) (a : α) : + set.range (foldl f a) = set.range (foldl g a) := +(foldl_range_subset_of_range_subset hfg.le a).antisymm (foldl_range_subset_of_range_subset hfg.ge a) + end list From 825edd3cd735e87495b0c2a2114fc3929eefce41 Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Fri, 27 Jan 2023 01:10:41 +0000 Subject: [PATCH 0369/1291] chore(ring_theory): protect `algebra.is_integral` (#18178) Co-authored-by: Junyan Xu --- src/field_theory/intermediate_field.lean | 2 +- .../minpoly/is_integrally_closed.lean | 8 ++++---- src/ring_theory/adjoin/power_basis.lean | 10 +++++----- src/ring_theory/discriminant.lean | 2 -- src/ring_theory/integral_closure.lean | 17 +++++++++-------- src/ring_theory/norm.lean | 12 ++++++------ src/ring_theory/trace.lean | 6 +++--- 7 files changed, 28 insertions(+), 29 deletions(-) diff --git a/src/field_theory/intermediate_field.lean b/src/field_theory/intermediate_field.lean index 8d184914aea9c..204b1b4415602 100644 --- a/src/field_theory/intermediate_field.lean +++ b/src/field_theory/intermediate_field.lean @@ -360,7 +360,7 @@ begin end lemma coe_is_integral_iff {R : Type*} [comm_ring R] [algebra R K] [algebra R L] - [is_scalar_tower R K L] {x : S} : is_integral R (x : L) ↔ _root_.is_integral R x := + [is_scalar_tower R K L] {x : S} : is_integral R (x : L) ↔ is_integral R x := begin refine ⟨λ h, _, λ h, _⟩, { obtain ⟨P, hPmo, hProot⟩ := h, diff --git a/src/field_theory/minpoly/is_integrally_closed.lean b/src/field_theory/minpoly/is_integrally_closed.lean index 4f6466607268e..8351ee5ba343f 100644 --- a/src/field_theory/minpoly/is_integrally_closed.lean +++ b/src/field_theory/minpoly/is_integrally_closed.lean @@ -180,14 +180,14 @@ alg_equiv.of_bijective (minpoly.to_adjoin R x) /-- The `power_basis` of `adjoin R {x}` given by `x`. See `algebra.adjoin.power_basis` for a version over a field. -/ -@[simps] def _root_.algebra.adjoin.power_basis' (hx : _root_.is_integral R x) : - _root_.power_basis R (algebra.adjoin R ({x} : set S)) := +@[simps] def _root_.algebra.adjoin.power_basis' (hx : is_integral R x) : + power_basis R (algebra.adjoin R ({x} : set S)) := power_basis.map (adjoin_root.power_basis' (minpoly.monic hx)) (minpoly.equiv_adjoin hx) /-- The power basis given by `x` if `B.gen ∈ adjoin R {x}`. -/ -@[simps] noncomputable def _root_.power_basis.of_gen_mem_adjoin' (B : _root_.power_basis R S) +@[simps] noncomputable def _root_.power_basis.of_gen_mem_adjoin' (B : power_basis R S) (hint : is_integral R x) (hx : B.gen ∈ adjoin R ({x} : set S)) : - _root_.power_basis R S := + power_basis R S := (algebra.adjoin.power_basis' hint).map $ (subalgebra.equiv_of_eq _ _ $ power_basis.adjoin_eq_top_of_gen_mem_adjoin hx).trans subalgebra.top_equiv diff --git a/src/ring_theory/adjoin/power_basis.lean b/src/ring_theory/adjoin/power_basis.lean index c91fc07156929..c4f18ea931fec 100644 --- a/src/ring_theory/adjoin/power_basis.lean +++ b/src/ring_theory/adjoin/power_basis.lean @@ -27,11 +27,11 @@ open_locale big_operators /-- The elements `1, x, ..., x ^ (d - 1)` for a basis for the `K`-module `K[x]`, where `d` is the degree of the minimal polynomial of `x`. -/ -noncomputable def adjoin.power_basis_aux {x : S} (hx : _root_.is_integral K x) : +noncomputable def adjoin.power_basis_aux {x : S} (hx : is_integral K x) : basis (fin (minpoly K x).nat_degree) K (adjoin K ({x} : set S)) := begin have hST : function.injective (algebra_map (adjoin K ({x} : set S)) S) := subtype.coe_injective, - have hx' : _root_.is_integral K + have hx' : is_integral K (show adjoin K ({x} : set S), from ⟨x, subset_adjoin (set.mem_singleton x)⟩), { apply (is_integral_algebra_map_iff hST).mp, convert hx, @@ -55,7 +55,7 @@ end /-- The power basis `1, x, ..., x ^ (d - 1)` for `K[x]`, where `d` is the degree of the minimal polynomial of `x`. See `algebra.adjoin.power_basis'` for a version over a more general base ring. -/ -@[simps gen dim] noncomputable def adjoin.power_basis {x : S} (hx : _root_.is_integral K x) : +@[simps gen dim] noncomputable def adjoin.power_basis {x : S} (hx : is_integral K x) : power_basis K (adjoin K ({x} : set S)) := { gen := ⟨x, subset_adjoin (set.mem_singleton x)⟩, dim := (minpoly K x).nat_degree, @@ -69,7 +69,7 @@ open algebra /-- The power basis given by `x` if `B.gen ∈ adjoin K {x}`. See `power_basis.of_gen_mem_adjoin'` for a version over a more general base ring. -/ @[simps] noncomputable def power_basis.of_gen_mem_adjoin {x : S} (B : power_basis K S) - (hint : _root_.is_integral K x) (hx : B.gen ∈ adjoin K ({x} : set S)) : power_basis K S := + (hint : is_integral K x) (hx : B.gen ∈ adjoin K ({x} : set S)) : power_basis K S := (algebra.adjoin.power_basis hint).map $ (subalgebra.equiv_of_eq _ _ $ power_basis.adjoin_eq_top_of_gen_mem_adjoin hx).trans subalgebra.top_equiv @@ -173,7 +173,7 @@ if `minpoly K B.gen = (minpoly R B.gen).map (algebra_map R L)`. This is the case if `R` is a GCD domain and `K` is its fraction ring. -/ lemma to_matrix_is_integral {B B' : power_basis K S} {P : R[X]} (h : aeval B.gen P = B'.gen) (hB : is_integral R B.gen) (hmin : minpoly K B.gen = (minpoly R B.gen).map (algebra_map R K)) : - ∀ i j, _root_.is_integral R (B.basis.to_matrix B'.basis i j) := + ∀ i j, is_integral R (B.basis.to_matrix B'.basis i j) := begin intros i j, rw [B.basis.to_matrix_apply, B'.coe_basis], diff --git a/src/ring_theory/discriminant.lean b/src/ring_theory/discriminant.lean index 1d156e5f4cf03..7c47309444ba8 100644 --- a/src/ring_theory/discriminant.lean +++ b/src/ring_theory/discriminant.lean @@ -261,8 +261,6 @@ section integral variables {R : Type z} [comm_ring R] [algebra R K] [algebra R L] [is_scalar_tower R K L] -local notation `is_integral` := _root_.is_integral - /-- If `K` and `L` are fields and `is_scalar_tower R K L`, and `b : ι → L` satisfies ` ∀ i, is_integral R (b i)`, then `is_integral R (discr K b)`. -/ lemma discr_is_integral {b : ι → L} (h : ∀ i, is_integral R (b i)) : diff --git a/src/ring_theory/integral_closure.lean b/src/ring_theory/integral_closure.lean index 49b9575bfd544..712da198454d9 100644 --- a/src/ring_theory/integral_closure.lean +++ b/src/ring_theory/integral_closure.lean @@ -59,7 +59,7 @@ def is_integral (x : A) : Prop := variable (A) /-- An algebra is integral if every element of the extension is integral over the base ring -/ -def algebra.is_integral : Prop := +protected def algebra.is_integral : Prop := (algebra_map R A).is_integral variables {R A} @@ -865,7 +865,7 @@ variables [algebra R A] [is_scalar_tower R A B] /-- If A is an R-algebra all of whose elements are integral over R, and x is an element of an A-algebra that is integral over A, then x is integral over R.-/ -lemma is_integral_trans (A_int : is_integral R A) (x : B) (hx : is_integral A x) : +lemma is_integral_trans (A_int : algebra.is_integral R A) (x : B) (hx : is_integral A x) : is_integral R x := begin rcases hx with ⟨p, pmonic, hp⟩, @@ -883,7 +883,8 @@ end /-- If A is an R-algebra all of whose elements are integral over R, and B is an A-algebra all of whose elements are integral over A, then all elements of B are integral over R.-/ -lemma algebra.is_integral_trans (hA : is_integral R A) (hB : is_integral A B) : is_integral R B := +lemma algebra.is_integral_trans (hA : algebra.is_integral R A) (hB : algebra.is_integral A B) : + algebra.is_integral R B := λ x, is_integral_trans hA x (hB x) lemma ring_hom.is_integral_trans (hf : f.is_integral) (hg : g.is_integral) : @@ -895,8 +896,8 @@ lemma ring_hom.is_integral_trans (hf : f.is_integral) (hg : g.is_integral) : lemma ring_hom.is_integral_of_surjective (hf : function.surjective f) : f.is_integral := λ x, (hf x).rec_on (λ y hy, (hy ▸ f.is_integral_map : f.is_integral_elem x)) -lemma is_integral_of_surjective (h : function.surjective (algebra_map R A)) : is_integral R A := -(algebra_map R A).is_integral_of_surjective h +lemma is_integral_of_surjective (h : function.surjective (algebra_map R A)) : + algebra.is_integral R A := (algebra_map R A).is_integral_of_surjective h /-- If `R → A → B` is an algebra tower with `A → B` injective, then if the entire tower is an integral extension so is `R → A` -/ @@ -949,8 +950,8 @@ begin simpa only [hom_eval₂, eval₂_map] using congr_arg (ideal.quotient.mk I) hpx end -lemma is_integral_quotient_of_is_integral {I : ideal A} (hRA : is_integral R A) : - is_integral (R ⧸ I.comap (algebra_map R A)) (A ⧸ I) := +lemma is_integral_quotient_of_is_integral {I : ideal A} (hRA : algebra.is_integral R A) : + algebra.is_integral (R ⧸ I.comap (algebra_map R A)) (A ⧸ I) := (algebra_map R A).is_integral_quotient_of_is_integral hRA lemma is_integral_quotient_map_iff {I : ideal S} : @@ -967,7 +968,7 @@ end /-- If the integral extension `R → S` is injective, and `S` is a field, then `R` is also a field. -/ lemma is_field_of_is_integral_of_is_field {R S : Type*} [comm_ring R] [nontrivial R] [comm_ring S] [is_domain S] - [algebra R S] (H : is_integral R S) (hRS : function.injective (algebra_map R S)) + [algebra R S] (H : algebra.is_integral R S) (hRS : function.injective (algebra_map R S)) (hS : is_field S) : is_field R := begin refine ⟨⟨0, 1, zero_ne_one⟩, mul_comm, λ a ha, _⟩, diff --git a/src/ring_theory/norm.lean b/src/ring_theory/norm.lean index f752917f45b59..ec9bca31b5aac 100644 --- a/src/ring_theory/norm.lean +++ b/src/ring_theory/norm.lean @@ -191,7 +191,7 @@ variable {K} section intermediate_field lemma _root_.intermediate_field.adjoin_simple.norm_gen_eq_one {x : L} - (hx : ¬_root_.is_integral K x) : norm K (adjoin_simple.gen K x) = 1 := + (hx : ¬is_integral K x) : norm K (adjoin_simple.gen K x) = 1 := begin rw [norm_eq_one_of_not_exists_basis], contrapose! hx, @@ -207,9 +207,9 @@ lemma _root_.intermediate_field.adjoin_simple.norm_gen_eq_prod_roots (x : L) ((minpoly K x).map (algebra_map K F)).roots.prod := begin have injKxL := (algebra_map K⟮x⟯ L).injective, - by_cases hx : _root_.is_integral K x, swap, + by_cases hx : is_integral K x, swap, { simp [minpoly.eq_zero hx, intermediate_field.adjoin_simple.norm_gen_eq_one hx] }, - have hx' : _root_.is_integral K (adjoin_simple.gen K x), + have hx' : is_integral K (adjoin_simple.gen K x), { rwa [← is_integral_algebra_map_iff injKxL, adjoin_simple.algebra_map_gen], apply_instance }, rw [← adjoin.power_basis_gen hx, power_basis.norm_gen_eq_prod_roots]; @@ -297,10 +297,10 @@ begin end lemma is_integral_norm [algebra R L] [algebra R K] [is_scalar_tower R K L] - [is_separable K L] [finite_dimensional K L] {x : L} (hx : _root_.is_integral R x) : - _root_.is_integral R (norm K x) := + [is_separable K L] [finite_dimensional K L] {x : L} (hx : is_integral R x) : + is_integral R (norm K x) := begin - have hx' : _root_.is_integral K x := is_integral_of_is_scalar_tower hx, + have hx' : is_integral K x := is_integral_of_is_scalar_tower hx, rw [← is_integral_algebra_map_iff (algebra_map K (algebraic_closure L)).injective, norm_eq_prod_roots], { refine (is_integral.multiset_prod (λ y hy, _)).pow _, diff --git a/src/ring_theory/trace.lean b/src/ring_theory/trace.lean index 8aa1462475762..91403bc5ddab4 100644 --- a/src/ring_theory/trace.lean +++ b/src/ring_theory/trace.lean @@ -283,10 +283,10 @@ variables [algebra R L] [algebra L F] [algebra R F] [is_scalar_tower R L F] open polynomial -lemma algebra.is_integral_trace [finite_dimensional L F] {x : F} (hx : _root_.is_integral R x) : - _root_.is_integral R (algebra.trace L F x) := +lemma algebra.is_integral_trace [finite_dimensional L F] {x : F} (hx : is_integral R x) : + is_integral R (algebra.trace L F x) := begin - have hx' : _root_.is_integral L x := is_integral_of_is_scalar_tower hx, + have hx' : is_integral L x := is_integral_of_is_scalar_tower hx, rw [← is_integral_algebra_map_iff (algebra_map L (algebraic_closure F)).injective, trace_eq_sum_roots], { refine (is_integral.multiset_sum _).nsmul _, From a1ddb5513dce7f8f4b9961549062e485aafb2c36 Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Fri, 27 Jan 2023 10:42:15 +0000 Subject: [PATCH 0370/1291] feat(analysis/special_functions/trigonometric): Euler's infinite product for sine (part a) (#18282) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This adds the first half of the proof of Euler's infinite product, writing `sin (π * z) / (π * z)` as the first n terms of the product times a ratio of cosine integrals. (The second half will be to show that this ratio of cosine integrals tends to 1). --- docs/references.bib | 15 ++ .../trigonometric/euler_sine_prod.lean | 255 ++++++++++++++++++ 2 files changed, 270 insertions(+) create mode 100644 src/analysis/special_functions/trigonometric/euler_sine_prod.lean diff --git a/docs/references.bib b/docs/references.bib index 3d2220c570725..d5afefcaf7e02 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -1872,6 +1872,21 @@ @Book{ rudin2006real isbn = {0-07-100276-6} } +@Article{ salwinski2018, + author = {Salwinski, David}, + title = {Euler's sine product formula: an elementary proof}, + journal = {College Math. J.}, + fjournal = {The College Mathematics Journal}, + volume = {49}, + year = {2018}, + number = {2}, + pages = {126--135}, + issn = {0746-8342}, + mrclass = {26A06 (00A05)}, + mrnumber = {3766700}, + doi = {10.1080/07468342.2018.1419703}, +} + @Book{ samuel1967, author = {Samuel, Pierre}, title = {Th\'{e}orie alg\'{e}brique des nombres}, diff --git a/src/analysis/special_functions/trigonometric/euler_sine_prod.lean b/src/analysis/special_functions/trigonometric/euler_sine_prod.lean new file mode 100644 index 0000000000000..8563d8909fec5 --- /dev/null +++ b/src/analysis/special_functions/trigonometric/euler_sine_prod.lean @@ -0,0 +1,255 @@ +/- +Copyright (c) 2023 David Loeffler. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: David Loeffler +-/ + +import analysis.special_functions.integrals +import analysis.special_functions.trigonometric.bounds +import data.real.pi.wallis + +/-! # Euler's infinite product for the sine function + +This file proves the infinite product formula + +$$ \sin \pi z = \pi z \prod_{n = 1}^\infty \left(1 - \frac{z ^ 2}{n ^ 2}\right) $$ + +for any real or complex `z`. Our proof closely follows the article +[Salwinski, *Euler's Sine Product Formula: An Elementary Proof*][salwinski2018]: the basic strategy +is to prove a recurrence relation for the integrals `∫ x in 0..π/2, cos 2 z x * cos x ^ (2 * n)`, +generalising the arguments used to prove Wallis' limit formula for `π`. +-/ + +open_locale real topological_space big_operators +open real set filter interval_integral measure_theory.measure_space + +namespace euler_sine + +section integral_recursion + +/-! ## Recursion formula for the integral of `cos (2 * z * x) * cos x ^ n` + +We evaluate the integral of `cos (2 * z * x) * cos x ^ n`, for any complex `z` and even integers +`n`, via repeated integration by parts. -/ + +variables {z : ℂ} {n : ℕ} + +lemma antideriv_cos_comp_const_mul (hz : z ≠ 0) (x : ℝ) : + has_deriv_at (λ y:ℝ, complex.sin (2 * z * y) / (2 * z)) (complex.cos (2 * z * x)) x := +begin + have a : has_deriv_at _ _ ↑x := has_deriv_at_mul_const _, + have b : has_deriv_at (λ (y : ℂ), complex.sin (y * (2 * z))) _ ↑x := + has_deriv_at.comp x (complex.has_deriv_at_sin (x * (2 * z))) a, + convert (b.comp_of_real).div_const (2 * z), + { ext1 x, rw mul_comm _ (2 * z) }, + { field_simp, rw mul_comm _ (2 * z) }, +end + +lemma antideriv_sin_comp_const_mul (hz : z ≠ 0) (x : ℝ) : + has_deriv_at (λ y:ℝ, -complex.cos (2 * z * y) / (2 * z)) (complex.sin (2 * z * x)) x := +begin + have a : has_deriv_at _ _ ↑x := has_deriv_at_mul_const _, + have b : has_deriv_at (λ (y : ℂ), complex.cos (y * (2 * z))) _ ↑x := + has_deriv_at.comp x (complex.has_deriv_at_cos (x * (2 * z))) a, + convert ((b.comp_of_real).div_const (2 * z)).neg, + { ext1 x, rw mul_comm _ (2 * z), field_simp }, + { field_simp, rw mul_comm _ (2 * z) }, +end + +lemma integral_cos_mul_cos_pow_aux (hn : 2 ≤ n) (hz : z ≠ 0): + (∫ x:ℝ in 0..π/2, complex.cos (2 * z * x) * cos x ^ n) = + n / (2 * z) * ∫ x:ℝ in 0..π/2, complex.sin (2 * z * x) * sin x * cos x ^ (n - 1) := +begin + have der1 : ∀ (x : ℝ), (x ∈ uIcc 0 (π/2)) → has_deriv_at (λ y, (↑(cos y)) ^ n : ℝ → ℂ) + (-n * sin x * cos x ^ (n - 1)) x, + { intros x hx, + have b : has_deriv_at (λ y, ↑(cos y) : ℝ → ℂ) (-sin x) x, + by simpa using (has_deriv_at_cos x).of_real_comp, + convert has_deriv_at.comp x (has_deriv_at_pow _ _) b using 1, + ring, }, + convert integral_mul_deriv_eq_deriv_mul der1 (λ x hx, antideriv_cos_comp_const_mul hz x) _ _, + { ext1 x, rw mul_comm }, + { rw [complex.of_real_zero, mul_zero, complex.sin_zero, zero_div, mul_zero, sub_zero, + cos_pi_div_two, complex.of_real_zero, zero_pow (by positivity : 0 < n), zero_mul, zero_sub, + ←integral_neg, ←integral_const_mul], + refine integral_congr (λ x hx, _), + field_simp, ring }, + { apply continuous.interval_integrable, + exact (continuous_const.mul (complex.continuous_of_real.comp continuous_sin)).mul + ((complex.continuous_of_real.comp continuous_cos).pow (n - 1)) }, + { apply continuous.interval_integrable, + exact complex.continuous_cos.comp (continuous_const.mul complex.continuous_of_real) } +end + +lemma integral_sin_mul_sin_mul_cos_pow_eq (hn : 2 ≤ n) (hz : z ≠ 0) : + ∫ x:ℝ in 0..π/2, complex.sin (2 * z * x) * sin x * cos x ^ (n - 1) = + n / (2 * z) * (∫ x:ℝ in 0..π/2, complex.cos (2 * z * x) * cos x ^ n) - + (n - 1) / (2 * z) * (∫ x:ℝ in 0..π/2, complex.cos (2 * z * x) * cos x ^ (n - 2)) := +begin + have der1 : ∀ (x : ℝ), (x ∈ uIcc 0 (π/2)) → + has_deriv_at (λ y, (sin y) * (cos y) ^ (n - 1) : ℝ → ℂ) + (cos x ^ n - (n - 1) * sin x ^ 2 * cos x ^ (n - 2)) x, + { intros x hx, + have c := has_deriv_at.comp (x:ℂ) (has_deriv_at_pow (n - 1) _) (complex.has_deriv_at_cos x), + convert ((complex.has_deriv_at_sin x).mul c).comp_of_real using 1, + { ext1 y, simp only [complex.of_real_sin, complex.of_real_cos] }, + { simp only [complex.of_real_cos, complex.of_real_sin], + rw [mul_neg, mul_neg, ←sub_eq_add_neg, function.comp_app], + congr' 1, + { rw [←pow_succ, nat.sub_add_cancel (by linarith : 1 ≤ n)] }, + { have : ((n - 1 : ℕ) : ℂ) = (n:ℂ) - 1, + { rw [nat.cast_sub (one_le_two.trans hn), nat.cast_one] }, + rw [nat.sub_sub, this], + ring } } }, + convert integral_mul_deriv_eq_deriv_mul der1 (λ x hx, antideriv_sin_comp_const_mul hz x) _ _ + using 1, + { refine integral_congr (λ x hx, _), + ring_nf }, + { -- now a tedious rearrangement of terms + -- gather into a single integral, and deal with continuity subgoals: + rw [sin_zero, cos_pi_div_two, complex.of_real_zero, zero_pow, zero_mul, mul_zero, zero_mul, + zero_mul, sub_zero, zero_sub, ←integral_neg, ←integral_const_mul, ←integral_const_mul, + ←integral_sub], + rotate, + { apply continuous.interval_integrable, + exact continuous_const.mul ((complex.continuous_cos.comp (continuous_const.mul + complex.continuous_of_real)).mul ((complex.continuous_of_real.comp + continuous_cos).pow n)) }, + { apply continuous.interval_integrable, + exact continuous_const.mul + ((complex.continuous_cos.comp (continuous_const.mul complex.continuous_of_real)).mul + ((complex.continuous_of_real.comp continuous_cos).pow (n - 2))), }, + { apply nat.sub_pos_of_lt, exact one_lt_two.trans_le hn }, + refine integral_congr (λ x hx, _), + dsimp only, + -- get rid of real trig functions and divions by 2 * z: + rw [complex.of_real_cos, complex.of_real_sin, complex.sin_sq, ←mul_div_right_comm, + ←mul_div_right_comm, ←sub_div, mul_div, ←neg_div], + congr' 1, + have : complex.cos ↑x ^ n = complex.cos ↑x ^ (n - 2) * complex.cos ↑x ^ 2, + { conv_lhs { rw [←nat.sub_add_cancel hn, pow_add] } }, + rw this, + ring }, + { apply continuous.interval_integrable, + exact ((complex.continuous_of_real.comp continuous_cos).pow n).sub + ((continuous_const.mul ((complex.continuous_of_real.comp continuous_sin).pow 2)).mul + ((complex.continuous_of_real.comp continuous_cos).pow (n - 2))) }, + { apply continuous.interval_integrable, + exact complex.continuous_sin.comp (continuous_const.mul complex.continuous_of_real) }, +end + +/-- Note this also holds for `z = 0`, but we do not need this case for `sin_pi_mul_eq`. -/ +lemma integral_cos_mul_cos_pow (hn : 2 ≤ n) (hz : z ≠ 0) : + (1 - 4 * z ^ 2 / n ^ 2) * (∫ x:ℝ in 0..π/2, complex.cos (2 * z * x) * cos x ^ n) = + (n - 1 : ℂ) / n * ∫ x:ℝ in 0..π/2, complex.cos (2 * z * x) * cos x ^ (n - 2) := +begin + have nne : (n : ℂ) ≠ 0, + { contrapose! hn, rw nat.cast_eq_zero at hn, rw hn, exact zero_lt_two }, + have := integral_cos_mul_cos_pow_aux hn hz, + rw [integral_sin_mul_sin_mul_cos_pow_eq hn hz, sub_eq_neg_add, mul_add, ←sub_eq_iff_eq_add] + at this, + convert congr_arg (λ u:ℂ, -u * (2 * z) ^ 2 / n ^ 2) this using 1; + { field_simp, ring }, +end + +/-- Note this also holds for `z = 0`, but we do not need this case for `sin_pi_mul_eq`. -/ +lemma integral_cos_mul_cos_pow_even (n : ℕ) (hz : z ≠ 0) : + (1 - z ^ 2 / (n + 1) ^ 2) * (∫ x:ℝ in 0..π/2, complex.cos (2 * z * x) * cos x ^ (2 * n + 2)) = + (2 * n + 1 : ℂ) / (2 * n + 2) * ∫ x:ℝ in 0..π/2, complex.cos (2 * z * x) * cos x ^ (2 * n) := +begin + convert integral_cos_mul_cos_pow (by linarith : 2 ≤ 2 * n + 2) hz using 3, + { simp only [nat.cast_add, nat.cast_mul, nat.cast_two], + nth_rewrite_rhs 2 ←mul_one (2:ℂ), + rw [←mul_add, mul_pow, ←div_div], + ring }, + { push_cast, ring }, + { push_cast, ring }, +end + +/-- Relate the integral `cos x ^ n` over `[0, π/2]` to the integral of `sin x ^ n` over `[0, π]`, +which is studied in `data.real.pi.wallis` and other places. -/ +lemma integral_cos_pow_eq (n : ℕ) : + (∫ (x:ℝ) in 0..π/2, cos x ^ n) = 1 / 2 * (∫ (x:ℝ) in 0..π, (sin x) ^ n) := +begin + rw [mul_comm (1/2 : ℝ), ←div_eq_iff (one_div_ne_zero (two_ne_zero' ℝ)), ←div_mul, div_one, + mul_two], + have L : interval_integrable _ volume 0 (π / 2) := (continuous_sin.pow n).interval_integrable _ _, + have R : interval_integrable _ volume (π / 2) π := (continuous_sin.pow n).interval_integrable _ _, + rw ←integral_add_adjacent_intervals L R, + congr' 1, + { nth_rewrite 0 (by ring : 0 = π/2 - π/2), + nth_rewrite 2 (by ring : π/2 = π/2 - 0), + rw ←integral_comp_sub_left, + refine integral_congr (λ x _, _), + dsimp only, + rw cos_pi_div_two_sub }, + { nth_rewrite 2 (by ring : π = π/2 + π/2), + nth_rewrite 1 (by ring : π/2 = 0 + π/2), + rw ←integral_comp_add_right, + refine integral_congr (λ x _, _), + dsimp only, + rw sin_add_pi_div_two }, +end + +lemma integral_cos_pow_pos (n : ℕ) : 0 < (∫ (x:ℝ) in 0..π/2, cos x ^ n) := +(integral_cos_pow_eq n).symm ▸ (mul_pos one_half_pos (integral_sin_pow_pos _)) + +/-- Finite form of Euler's sine product, with remainder term expressed as a ratio of cosine +integrals. -/ +lemma sin_pi_mul_eq (z : ℂ) (n : ℕ) : + complex.sin (π * z) = π * z * (∏ j in finset.range n, (1 - z ^ 2 / (j + 1) ^ 2)) * + (∫ x in 0..π/2, complex.cos (2 * z * x) * cos x ^ (2 * n)) / ↑∫ x in 0..π/2, cos x ^ (2 * n) := +begin + rcases eq_or_ne z 0 with rfl | hz, + { simp }, + induction n with n hn, + { simp_rw [mul_zero, pow_zero, mul_one, finset.prod_range_zero, mul_one, integral_one, sub_zero], + rw [integral_cos_mul_complex (mul_ne_zero two_ne_zero hz), complex.of_real_zero, mul_zero, + complex.sin_zero, zero_div, sub_zero, + (by { push_cast, field_simp, ring } : 2 * z * ↑(π / 2) = π * z)], + field_simp [complex.of_real_ne_zero.mpr pi_pos.ne'], + ring }, + { rw [hn, finset.prod_range_succ], + set A := ∏ j in finset.range n, (1 - z ^ 2 / (j + 1) ^ 2), + set B := ∫ x:ℝ in 0..π/2, complex.cos (2 * z * x) * cos x ^ (2 * n), + set C := ∫ x:ℝ in 0..π/2, cos x ^ (2 * n), + have aux' : 2 * n.succ = 2 * n + 2, + { rw [nat.succ_eq_add_one, mul_add, mul_one], }, + have : ∫ x:ℝ in 0..π/2, cos x ^ (2 * n.succ) = (2 * (n:ℝ) + 1) / (2 * n + 2) * C, + { rw integral_cos_pow_eq, + dsimp only [C], + rw [integral_cos_pow_eq, aux', integral_sin_pow, sin_zero, sin_pi, pow_succ, zero_mul, + zero_mul, zero_mul, sub_zero, zero_div, zero_add, ←mul_assoc, ←mul_assoc, + mul_comm (1 / 2 : ℝ) _, nat.cast_mul, nat.cast_bit0, nat.cast_one] }, + rw this, + change ↑π * z * A * B / ↑C = + (↑π * z * (A * (1 - z ^ 2 / (↑n + 1) ^ 2)) * + ∫ (x : ℝ) in 0..π / 2, complex.cos (2 * z * ↑x) * ↑(cos x) ^ (2 * n.succ)) / + ↑((2 * ↑n + 1) / (2 * ↑n + 2) * C), + have : ↑π * z * (A * (1 - z ^ 2 / (↑n + 1) ^ 2)) * + ∫ (x : ℝ) in 0..π / 2, complex.cos (2 * z * ↑x) * ↑(cos x) ^ (2 * n.succ) + = ↑π * z * A * ((1 - z ^ 2 / (↑n.succ) ^ 2) * + ∫ (x : ℝ) in 0..π / 2, complex.cos (2 * z * ↑x) * ↑(cos x) ^ (2 * n.succ)), + { nth_rewrite_rhs 0 nat.succ_eq_add_one, + rw nat.cast_add_one, + ring }, + rw this, + suffices : (1 - z ^ 2 / ↑(n.succ) ^ 2) * + ∫ (x : ℝ) in 0..π / 2, complex.cos (2 * z * ↑x) * ↑(cos x) ^ (2 * n.succ) = + (2 * n + 1) / (2 * n + 2) * B, + { rw [this, complex.of_real_mul, complex.of_real_div], + have : (C:ℂ) ≠ 0 := complex.of_real_ne_zero.mpr (integral_cos_pow_pos _).ne', + have : 2 * (n:ℂ) + 1 ≠ 0, + { convert (nat.cast_add_one_ne_zero (2 * n) : (↑(2 * n) + 1 : ℂ) ≠ 0), + simp }, + have : 2 * (n:ℂ) + 2 ≠ 0, + { convert (nat.cast_add_one_ne_zero (2 * n + 1) : (↑(2 * n + 1) + 1 : ℂ) ≠ 0) using 1, + push_cast, ring }, + field_simp, ring }, + convert integral_cos_mul_cos_pow_even n hz, + rw nat.cast_succ } +end + +end integral_recursion + +end euler_sine From fedcb65cb16a854dd9ffffc3fb62b1d453b12678 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Fri, 27 Jan 2023 11:49:06 +0000 Subject: [PATCH 0371/1291] feat(set_theory/zfc/basic): more trivial results (#18294) This PR does multiple very simple things at once. Here's a rundown. - Golf `Set.eq_empty`. - Add `Set.eq_empty_or_nonempty`, `Set.insert_nonempty`, `Set.singleton_nonempty`, `Class.mem_def`, `Class.not_empty_hom`, `Class.univ_hom`, `Class.empty_Cong_to_Class`, `Class.empty_Class_to_Cong`. - Tag `Class.mem_univ` as `simp`. - Move `Set.to_set_sUnion` and `Set.sUnion_empty`, so that the theorems on singleton injectivity aren't interspersed with the union results. --- src/set_theory/zfc/basic.lean | 35 ++++++++++++++++++++++++++--------- 1 file changed, 26 insertions(+), 9 deletions(-) diff --git a/src/set_theory/zfc/basic.lean b/src/set_theory/zfc/basic.lean index bdb8336eaa245..d5b8d9422ab14 100644 --- a/src/set_theory/zfc/basic.lean +++ b/src/set_theory/zfc/basic.lean @@ -549,9 +549,10 @@ begin exact ⟨a, h⟩ end -theorem eq_empty (x : Set.{u}) : x = ∅ ↔ ∀ y : Set.{u}, y ∉ x := -⟨λ h y, (h.symm ▸ not_mem_empty y), -λ h, ext (λ y, ⟨λ yx, absurd yx (h y), λ y0, absurd y0 (not_mem_empty _)⟩)⟩ +theorem eq_empty (x : Set.{u}) : x = ∅ ↔ ∀ y : Set.{u}, y ∉ x := by { rw ext_iff, simp } + +theorem eq_empty_or_nonempty (u : Set) : u = ∅ ∨ u.nonempty := +by { rw [eq_empty, ←not_exists], apply em' } /-- `insert x y` is the set `{x} ∪ y` -/ protected def insert : Set → Set → Set := @@ -594,6 +595,10 @@ iff.trans mem_insert_iff ⟨λ o, or.rec (λ h, h) (λ n, absurd n (not_mem_empt @[simp] theorem to_set_singleton (x : Set) : ({x} : Set).to_set = {x} := by { ext, simp } +theorem insert_nonempty (u v : Set) : (insert u v).nonempty := ⟨u, mem_insert u v⟩ + +theorem singleton_nonempty (u : Set) : Set.nonempty {u} := insert_nonempty u ∅ + @[simp] theorem mem_pair {x y z : Set.{u}} : x ∈ ({y, z} : Set) ↔ x = y ∨ x = z := iff.trans mem_insert_iff $ or_congr iff.rfl mem_singleton @@ -669,19 +674,19 @@ quotient.induction_on₂ x y (λ x y, iff.trans mem_sUnion theorem mem_sUnion_of_mem {x y z : Set} (hy : y ∈ z) (hz : z ∈ x) : y ∈ ⋃₀ x := mem_sUnion.2 ⟨z, hz, hy⟩ +@[simp] theorem sUnion_empty : ⋃₀ (∅ : Set.{u}) = ∅ := by { ext, simp } + @[simp] theorem sUnion_singleton {x : Set.{u}} : ⋃₀ ({x} : Set) = x := ext $ λ y, by simp_rw [mem_sUnion, exists_prop, mem_singleton, exists_eq_left] +@[simp] theorem to_set_sUnion (x : Set.{u}) : (⋃₀ x).to_set = ⋃₀ (to_set '' x.to_set) := +by { ext, simp } + theorem singleton_injective : function.injective (@singleton Set Set _) := λ x y H, let this := congr_arg sUnion H in by rwa [sUnion_singleton, sUnion_singleton] at this @[simp] theorem singleton_inj {x y : Set} : ({x} : Set) = {y} ↔ x = y := singleton_injective.eq_iff -@[simp] theorem to_set_sUnion (x : Set.{u}) : (⋃₀ x).to_set = ⋃₀ (to_set '' x.to_set) := -by { ext, simp } - -@[simp] theorem sUnion_empty : ⋃₀ (∅ : Set.{u}) = ∅ := by { ext, simp } - /-- The binary union operation -/ protected def union (x y : Set.{u}) : Set.{u} := ⋃₀ {x, y} @@ -879,11 +884,17 @@ def to_Set (B : Class.{u}) (A : Class.{u}) : Prop := ∃ x, ↑x = A ∧ B x protected def mem (A B : Class.{u}) : Prop := to_Set.{u} B A instance : has_mem Class Class := ⟨Class.mem⟩ +theorem mem_def (A B : Class.{u}) : A ∈ B ↔ ∃ x, ↑x = A ∧ B x := iff.rfl + @[simp] theorem not_mem_empty (x : Class.{u}) : x ∉ (∅ : Class.{u}) := λ ⟨_, _, h⟩, h -theorem mem_univ {A : Class.{u}} : A ∈ univ.{u} ↔ ∃ x : Set.{u}, ↑x = A := +@[simp] theorem not_empty_hom (x : Set.{u}) : ¬ (∅ : Class.{u}) x := id + +@[simp] theorem mem_univ {A : Class.{u}} : A ∈ univ.{u} ↔ ∃ x : Set.{u}, ↑x = A := exists_congr $ λ x, and_true _ +@[simp] theorem mem_univ_hom (x : Set.{u}) : univ.{u} x := trivial + theorem mem_wf : @well_founded Class.{u} (∈) := ⟨begin have H : ∀ x : Set.{u}, @acc Class.{u} (∈) ↑x, @@ -910,9 +921,15 @@ theorem univ_not_mem_univ : univ ∉ univ := mem_irrefl _ /-- Convert a conglomerate (a collection of classes) into a class -/ def Cong_to_Class (x : set Class.{u}) : Class.{u} := {y | ↑y ∈ x} +@[simp] theorem Cong_to_Class_empty : Cong_to_Class ∅ = ∅ := +by { ext, simp [Cong_to_Class] } + /-- Convert a class into a conglomerate (a collection of classes) -/ def Class_to_Cong (x : Class.{u}) : set Class.{u} := {y | y ∈ x} +@[simp] theorem Class_to_Cong_empty : Class_to_Cong ∅ = ∅ := +by { ext, simp [Class_to_Cong] } + /-- The power class of a class is the class of all subclasses that are ZFC sets -/ def powerset (x : Class) : Class := Cong_to_Class (set.powerset x) From a11f9106a169dd302a285019e5165f8ab32ff433 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Fri, 27 Jan 2023 14:19:56 +0000 Subject: [PATCH 0372/1291] chore(*): add mathlib4 synchronization comments (#18314) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.group_ring_action.subobjects` * `algebra.quandle` * `data.bitvec.core` * `data.bool.count` * `data.enat.basic` * `data.fin.interval` * `data.fin.tuple.monotone` * `data.finset.locally_finite` * `data.finset.pi_induction` * `data.fintype.fin` * `data.int.log` * `data.int.parity` * `data.nat.interval` * `data.nat.parity` * `data.pfunctor.univariate.basic` * `deprecated.submonoid` * `group_theory.archimedean` * `group_theory.group_action.basic` * `group_theory.subgroup.actions` * `group_theory.subgroup.basic` * `group_theory.subgroup.mul_opposite` * `group_theory.subgroup.zpowers` * `order.filter.interval` * `order.locally_finite` * `order.partial_sups` * `topology.basic` * `topology.bornology.basic` * `topology.bornology.constructions` --- src/algebra/group_ring_action/subobjects.lean | 3 +++ src/algebra/quandle.lean | 3 +++ src/data/bitvec/core.lean | 3 +++ src/data/bool/count.lean | 3 +++ src/data/enat/basic.lean | 3 +++ src/data/fin/interval.lean | 3 +++ src/data/fin/tuple/monotone.lean | 3 +++ src/data/finset/locally_finite.lean | 3 +++ src/data/finset/pi_induction.lean | 3 +++ src/data/fintype/fin.lean | 3 +++ src/data/int/log.lean | 3 +++ src/data/int/parity.lean | 3 +++ src/data/nat/interval.lean | 3 +++ src/data/nat/parity.lean | 3 +++ src/data/pfunctor/univariate/basic.lean | 3 +++ src/deprecated/submonoid.lean | 3 +++ src/group_theory/archimedean.lean | 3 +++ src/group_theory/group_action/basic.lean | 3 +++ src/group_theory/subgroup/actions.lean | 3 +++ src/group_theory/subgroup/basic.lean | 3 +++ src/group_theory/subgroup/mul_opposite.lean | 3 +++ src/group_theory/subgroup/zpowers.lean | 3 +++ src/order/filter/interval.lean | 3 +++ src/order/locally_finite.lean | 3 +++ src/order/partial_sups.lean | 3 +++ src/topology/basic.lean | 3 +++ src/topology/bornology/basic.lean | 3 +++ src/topology/bornology/constructions.lean | 3 +++ 28 files changed, 84 insertions(+) diff --git a/src/algebra/group_ring_action/subobjects.lean b/src/algebra/group_ring_action/subobjects.lean index 6857d2e8969a0..638436b5ccd3c 100644 --- a/src/algebra/group_ring_action/subobjects.lean +++ b/src/algebra/group_ring_action/subobjects.lean @@ -9,6 +9,9 @@ import group_theory.subgroup.basic /-! # Instances of `mul_semiring_action` for subobjects +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + These are defined in this file as `semiring`s are not available yet where `submonoid` and `subgroup` are defined. diff --git a/src/algebra/quandle.lean b/src/algebra/quandle.lean index 481d9047c8649..a1e9b87f20d72 100644 --- a/src/algebra/quandle.lean +++ b/src/algebra/quandle.lean @@ -11,6 +11,9 @@ import tactic.group /-! # Racks and Quandles +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines racks and quandles, algebraic structures for sets that bijectively act on themselves with a self-distributivity property. If `R` is a rack and `act : R → (R ≃ R)` is the self-action, diff --git a/src/data/bitvec/core.lean b/src/data/bitvec/core.lean index 8b3fbda636778..cc1f908c735d8 100644 --- a/src/data/bitvec/core.lean +++ b/src/data/bitvec/core.lean @@ -10,6 +10,9 @@ import data.nat.pow /-! # Basic operations on bitvectors +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This is a work-in-progress, and contains additions to other theories. This file was moved to mathlib from core Lean in the switch to Lean 3.20.0c. diff --git a/src/data/bool/count.lean b/src/data/bool/count.lean index 2ac772d0fb6fd..bfb0b4dc95a21 100644 --- a/src/data/bool/count.lean +++ b/src/data/bool/count.lean @@ -9,6 +9,9 @@ import data.list.chain /-! # List of booleans +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove lemmas about the number of `ff`s and `tt`s in a list of booleans. First we prove that the number of `ff`s plus the number of `tt` equals the length of the list. Then we prove that in a list with alternating `tt`s and `ff`s, the number of `tt`s differs from the number of diff --git a/src/data/enat/basic.lean b/src/data/enat/basic.lean index c3200de999f2a..b344502039c9d 100644 --- a/src/data/enat/basic.lean +++ b/src/data/enat/basic.lean @@ -11,6 +11,9 @@ import algebra.order.ring.with_top /-! # Definition and basic properties of extended natural numbers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define `enat` (notation: `ℕ∞`) to be `with_top ℕ` and prove some basic lemmas about this type. -/ diff --git a/src/data/fin/interval.lean b/src/data/fin/interval.lean index af0d1b267af69..41a0f2409e519 100644 --- a/src/data/fin/interval.lean +++ b/src/data/fin/interval.lean @@ -9,6 +9,9 @@ import data.finset.locally_finite /-! # Finite intervals in `fin n` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves that `fin n` is a `locally_finite_order` and calculates the cardinality of its intervals as finsets and fintypes. -/ diff --git a/src/data/fin/tuple/monotone.lean b/src/data/fin/tuple/monotone.lean index d5973fc02cc03..e67d043a65cb3 100644 --- a/src/data/fin/tuple/monotone.lean +++ b/src/data/fin/tuple/monotone.lean @@ -8,6 +8,9 @@ import data.fin.vec_notation /-! # Monotone finite sequences +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove `simp` lemmas that allow to simplify propositions like `monotone ![a, b, c]`. -/ diff --git a/src/data/finset/locally_finite.lean b/src/data/finset/locally_finite.lean index 6e3cf7f748223..c495c94620e52 100644 --- a/src/data/finset/locally_finite.lean +++ b/src/data/finset/locally_finite.lean @@ -9,6 +9,9 @@ import data.set.intervals.monoid /-! # Intervals as finsets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file provides basic results about all the `finset.Ixx`, which are defined in `order.locally_finite`. diff --git a/src/data/finset/pi_induction.lean b/src/data/finset/pi_induction.lean index 9cc1013b99ec7..646a65e5c0126 100644 --- a/src/data/finset/pi_induction.lean +++ b/src/data/finset/pi_induction.lean @@ -9,6 +9,9 @@ import data.finset.sigma /-! # Induction principles for `Π i, finset (α i)` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove a few induction principles for functions `Π i : ι, finset (α i)` defined on a finite type. diff --git a/src/data/fintype/fin.lean b/src/data/fintype/fin.lean index 1e2d862e57c98..7c834516819b4 100644 --- a/src/data/fintype/fin.lean +++ b/src/data/fintype/fin.lean @@ -8,6 +8,9 @@ import data.fin.interval /-! # The structure of `fintype (fin n)` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains some basic results about the `fintype` instance for `fin`, especially properties of `finset.univ : finset (fin n)`. -/ diff --git a/src/data/int/log.lean b/src/data/int/log.lean index 9beea3d18ed73..2ce1f2a995695 100644 --- a/src/data/int/log.lean +++ b/src/data/int/log.lean @@ -9,6 +9,9 @@ import data.nat.log /-! # Integer logarithms in a field with respect to a natural base +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines two `ℤ`-valued analogs of the logarithm of `r : R` with base `b : ℕ`: * `int.log b r`: Lower logarithm, or floor **log**. Greatest `k` such that `↑b^k ≤ r`. diff --git a/src/data/int/parity.lean b/src/data/int/parity.lean index a012ad282e206..945b5cec31bad 100644 --- a/src/data/int/parity.lean +++ b/src/data/int/parity.lean @@ -8,6 +8,9 @@ import data.nat.parity /-! # Parity of integers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains theorems about the `even` and `odd` predicates on the integers. ## Tags diff --git a/src/data/nat/interval.lean b/src/data/nat/interval.lean index 2b21651334411..d1a5a2e862753 100644 --- a/src/data/nat/interval.lean +++ b/src/data/nat/interval.lean @@ -8,6 +8,9 @@ import data.finset.locally_finite /-! # Finite intervals of naturals +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves that `ℕ` is a `locally_finite_order` and calculates the cardinality of its intervals as finsets and fintypes. diff --git a/src/data/nat/parity.lean b/src/data/nat/parity.lean index 2acd98c48dd03..6a0cf0a24bcbc 100644 --- a/src/data/nat/parity.lean +++ b/src/data/nat/parity.lean @@ -9,6 +9,9 @@ import algebra.parity /-! # Parity of natural numbers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains theorems about the `even` and `odd` predicates on the natural numbers. ## Tags diff --git a/src/data/pfunctor/univariate/basic.lean b/src/data/pfunctor/univariate/basic.lean index ba63040d31e2b..f969ae2c868ff 100644 --- a/src/data/pfunctor/univariate/basic.lean +++ b/src/data/pfunctor/univariate/basic.lean @@ -8,6 +8,9 @@ import data.W.basic /-! # Polynomial functors +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines polynomial functors and the W-type construction as a polynomial functor. (For the M-type construction, see pfunctor/M.lean.) diff --git a/src/deprecated/submonoid.lean b/src/deprecated/submonoid.lean index 7750164c2a044..d8887860703e0 100644 --- a/src/deprecated/submonoid.lean +++ b/src/deprecated/submonoid.lean @@ -10,6 +10,9 @@ import deprecated.group /-! # Unbundled submonoids (deprecated) +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file is deprecated, and is no longer imported by anything in mathlib other than other deprecated files, and test files. You should not need to import it. diff --git a/src/group_theory/archimedean.lean b/src/group_theory/archimedean.lean index d5732aa7bbd2b..7efb0f3ef1c51 100644 --- a/src/group_theory/archimedean.lean +++ b/src/group_theory/archimedean.lean @@ -9,6 +9,9 @@ import group_theory.subgroup.basic /-! # Archimedean groups +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves a few facts about ordered groups which satisfy the `archimedean` property, that is: `class archimedean (α) [ordered_add_comm_monoid α] : Prop :=` `(arch : ∀ (x : α) {y}, 0 < y → ∃ n : ℕ, x ≤ n • y)` diff --git a/src/group_theory/group_action/basic.lean b/src/group_theory/group_action/basic.lean index faec650307fa1..767fc4042cae6 100644 --- a/src/group_theory/group_action/basic.lean +++ b/src/group_theory/group_action/basic.lean @@ -13,6 +13,9 @@ import group_theory.subgroup.basic /-! # Basic properties of group actions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file primarily concerns itself with orbits, stabilizers, and other objects defined in terms of actions. Despite this file being called `basic`, low-level helper lemmas for algebraic manipulation of `•` belong elsewhere. diff --git a/src/group_theory/subgroup/actions.lean b/src/group_theory/subgroup/actions.lean index c3d8d0ad3fa79..53a3f2a0e46f0 100644 --- a/src/group_theory/subgroup/actions.lean +++ b/src/group_theory/subgroup/actions.lean @@ -9,6 +9,9 @@ import group_theory.subgroup.basic /-! # Actions by `subgroup`s +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + These are just copies of the definitions about `submonoid` starting from `submonoid.mul_action`. ## Tags diff --git a/src/group_theory/subgroup/basic.lean b/src/group_theory/subgroup/basic.lean index 8e4c2f27970bd..0b4b0623c8749 100644 --- a/src/group_theory/subgroup/basic.lean +++ b/src/group_theory/subgroup/basic.lean @@ -15,6 +15,9 @@ import tactic.apply_fun /-! # Subgroups +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines multiplicative and additive subgroups as an extension of submonoids, in a bundled form (unbundled subgroups are in `deprecated/subgroups.lean`). diff --git a/src/group_theory/subgroup/mul_opposite.lean b/src/group_theory/subgroup/mul_opposite.lean index 93f590238c684..acc7c3d591f76 100644 --- a/src/group_theory/subgroup/mul_opposite.lean +++ b/src/group_theory/subgroup/mul_opposite.lean @@ -9,6 +9,9 @@ import group_theory.subgroup.actions /-! # Mul-opposite subgroups +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Tags subgroup, subgroups diff --git a/src/group_theory/subgroup/zpowers.lean b/src/group_theory/subgroup/zpowers.lean index b7a89fb456edd..ca2ee8f68e936 100644 --- a/src/group_theory/subgroup/zpowers.lean +++ b/src/group_theory/subgroup/zpowers.lean @@ -9,6 +9,9 @@ import group_theory.subgroup.basic /-! # Subgroups generated by an element +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Tags subgroup, subgroups diff --git a/src/order/filter/interval.lean b/src/order/filter/interval.lean index ce637e7d62fdb..680da810a9f77 100644 --- a/src/order/filter/interval.lean +++ b/src/order/filter/interval.lean @@ -10,6 +10,9 @@ import order.filter.at_top_bot /-! # Convergence of intervals +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + If both `a` and `b` tend to some filter `l₁`, sometimes this implies that `Ixx a b` tends to `l₂.small_sets`, i.e., for any `s ∈ l₂` eventually `Ixx a b` becomes a subset of `s`. Here and below `Ixx` is one of `Icc`, `Ico`, `Ioc`, and `Ioo`. We define `filter.tendsto_Ixx_class Ixx l₁ l₂` diff --git a/src/order/locally_finite.lean b/src/order/locally_finite.lean index e1e3238b53a60..40c3ec825ac84 100644 --- a/src/order/locally_finite.lean +++ b/src/order/locally_finite.lean @@ -9,6 +9,9 @@ import data.set.intervals.unordered_interval /-! # Locally finite orders +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines locally finite orders. A locally finite order is an order for which all bounded intervals are finite. This allows to make diff --git a/src/order/partial_sups.lean b/src/order/partial_sups.lean index 41e97ca76bed3..d8b326407c8d4 100644 --- a/src/order/partial_sups.lean +++ b/src/order/partial_sups.lean @@ -10,6 +10,9 @@ import order.conditionally_complete_lattice.finset /-! # The monotone sequence of partial supremums of a sequence +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define `partial_sups : (ℕ → α) → ℕ →o α` inductively. For `f : ℕ → α`, `partial_sups f` is the sequence `f 0 `, `f 0 ⊔ f 1`, `f 0 ⊔ f 1 ⊔ f 2`, ... The point of this definition is that * it doesn't need a `⨆`, as opposed to `⨆ (i ≤ n), f i` (which also means the wrong thing on diff --git a/src/topology/basic.lean b/src/topology/basic.lean index 9b62d1a97598b..5b88faadd7d39 100644 --- a/src/topology/basic.lean +++ b/src/topology/basic.lean @@ -10,6 +10,9 @@ import order.filter.lift /-! # Basic theory of topological spaces. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The main definition is the type class `topological_space α` which endows a type `α` with a topology. Then `set α` gets predicates `is_open`, `is_closed` and functions `interior`, `closure` and `frontier`. Each point `x` of `α` gets a neighborhood filter `𝓝 x`. A filter `F` on `α` has diff --git a/src/topology/bornology/basic.lean b/src/topology/bornology/basic.lean index 80f7ae0d71c7a..6d518439fe531 100644 --- a/src/topology/bornology/basic.lean +++ b/src/topology/bornology/basic.lean @@ -8,6 +8,9 @@ import order.filter.cofinite /-! # Basic theory of bornology +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We develop the basic theory of bornologies. Instead of axiomatizing bounded sets and defining bornologies in terms of those, we recognize that the cobounded sets form a filter and define a bornology as a filter of cobounded sets which contains the cofinite filter. This allows us to make diff --git a/src/topology/bornology/constructions.lean b/src/topology/bornology/constructions.lean index f0eb1fc12d236..3b9212fa47b0f 100644 --- a/src/topology/bornology/constructions.lean +++ b/src/topology/bornology/constructions.lean @@ -8,6 +8,9 @@ import topology.bornology.basic /-! # Bornology structure on products and subtypes +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + 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. From 59cdeb0da2480abbc235b7e611ccd9a7e5603d7c Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Fri, 27 Jan 2023 17:10:19 +0000 Subject: [PATCH 0373/1291] feat(algebra/module/graded_module): define graded module (#14582) By imitating the current `graded_algebra`, this pr builds `graded_module` over some `graded algebra` Co-authored-by: Eric Wieser @eric-wieser - [x] depends on: #14626 - [x] depends on: #15654 Co-authored-by: Eric Wieser --- src/algebra/graded_mul_action.lean | 135 +++++++++++++++ src/algebra/module/graded_module.lean | 230 ++++++++++++++++++++++++++ 2 files changed, 365 insertions(+) create mode 100644 src/algebra/graded_mul_action.lean create mode 100644 src/algebra/module/graded_module.lean diff --git a/src/algebra/graded_mul_action.lean b/src/algebra/graded_mul_action.lean new file mode 100644 index 0000000000000..c70754ef199a2 --- /dev/null +++ b/src/algebra/graded_mul_action.lean @@ -0,0 +1,135 @@ +/- +Copyright (c) 2022 Jujian Zhang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jujian Zhang, Eric Wieser +-/ +import algebra.graded_monoid + +/-! +# Additively-graded multiplicative action structures + +This module provides a set of heterogeneous typeclasses for defining a multiplicative structure +over the sigma type `graded_monoid A` such that `(•) : A i → M j → M (i + j)`; that is to say, `A` +has an additively-graded multiplicative action on `M`. The typeclasses are: + +* `graded_monoid.ghas_smul A M` +* `graded_monoid.gmul_action A M` + +With the `sigma_graded` locale open, these respectively imbue: + +* `has_smul (graded_monoid A) (graded_monoid M)` +* `mul_action (graded_monoid A) (graded_monoid M)` + +For now, these typeclasses are primarily used in the construction of `direct_sum.gmodule.module` and +the rest of that file. + +## Internally graded multiplicative actions + +In addition to the above typeclasses, in the most frequent case when `A` is an indexed collection of +`set_like` subobjects (such as `add_submonoid`s, `add_subgroup`s, or `submodule`s), this file +provides the `Prop` typeclasses: + +* `set_like.has_graded_smul A M` (which provides the obvious `graded_monoid.ghas_smul A` instance) + +which provides the API lemma + +* `set_like.graded_smul_mem_graded` + +Note that there is no need for `set_like.graded_mul_action` or similar, as all the information it +would contain is already supplied by `has_graded_smul` when the objects within `A` and `M` have +a `mul_action` instance. + +## tags + +graded action +-/ + +set_option old_structure_cmd true + +variables {ι : Type*} + +namespace graded_monoid + +/-! ### Typeclasses -/ +section defs + +variables (A : ι → Type*) (M : ι → Type*) + +/-- A graded version of `has_smul`. Scalar multiplication combines grades additively, i.e. +if `a ∈ A i` and `m ∈ M j`, then `a • b` must be in `M (i + j)`-/ +class ghas_smul [has_add ι] := +(smul {i j} : A i → M j → M (i + j)) + +/-- A graded version of `has_mul.to_has_smul` -/ +instance ghas_mul.to_ghas_smul [has_add ι] [ghas_mul A] : ghas_smul A A := +{ smul := λ _ _, ghas_mul.mul } + +instance ghas_smul.to_has_smul [has_add ι] [ghas_smul A M] : + has_smul (graded_monoid A) (graded_monoid M) := +⟨λ (x : graded_monoid A) (y : graded_monoid M), ⟨_, ghas_smul.smul x.snd y.snd⟩⟩ + +lemma mk_smul_mk [has_add ι] [ghas_smul A M] {i j} (a : A i) (b : M j) : + mk i a • mk j b = mk (i + j) (ghas_smul.smul a b) := +rfl + +/-- A graded version of `mul_action`. -/ +class gmul_action [add_monoid ι] [gmonoid A] extends ghas_smul A M := +(one_smul (b : graded_monoid M) : (1 : graded_monoid A) • b = b) +(mul_smul (a a' : graded_monoid A) (b : graded_monoid M) : (a * a') • b = a • a' • b) + +/-- The graded version of `monoid.to_mul_action`. -/ +instance gmonoid.to_gmul_action [add_monoid ι] [gmonoid A] : + gmul_action A A := +{ one_smul := gmonoid.one_mul, + mul_smul := gmonoid.mul_assoc, + ..ghas_mul.to_ghas_smul _ } + +instance gmul_action.to_mul_action [add_monoid ι] [gmonoid A] [gmul_action A M] : + mul_action (graded_monoid A) (graded_monoid M) := +{ one_smul := gmul_action.one_smul, + mul_smul := gmul_action.mul_smul } + +end defs + +end graded_monoid + +/-! ### Shorthands for creating instance of the above typeclasses for collections of subobjects -/ + +section subobjects + +variables {R : Type*} + +/-- A version of `graded_monoid.ghas_smul` for internally graded objects. -/ +class set_like.has_graded_smul {S R N M : Type*} [set_like S R] [set_like N M] + [has_smul R M] [has_add ι] (A : ι → S) (B : ι → N) : Prop := +(smul_mem : ∀ ⦃i j : ι⦄ {ai bj}, ai ∈ A i → bj ∈ B j → ai • bj ∈ B (i + j)) + +instance set_like.ghas_smul {S R N M : Type*} [set_like S R] [set_like N M] + [has_smul R M] [has_add ι] (A : ι → S) (B : ι → N) [set_like.has_graded_smul A B] : + graded_monoid.ghas_smul (λ i, A i) (λ i, B i) := +{ smul := λ i j a b, ⟨(a : R) • b, set_like.has_graded_smul.smul_mem a.2 b.2⟩ } + +@[simp] lemma set_like.coe_ghas_smul {S R N M : Type*} [set_like S R] [set_like N M] + [has_smul R M] [has_add ι] (A : ι → S) (B : ι → N) [set_like.has_graded_smul A B] + {i j : ι} (x : A i) (y : B j) : + (@graded_monoid.ghas_smul.smul ι (λ i, A i) (λ i, B i) _ _ i j x y : M) = ((x : R) • y) := +rfl + +/-- Internally graded version of `has_mul.to_has_smul`. -/ +instance set_like.has_graded_mul.to_has_graded_smul [add_monoid ι] [monoid R] + {S : Type*} [set_like S R] (A : ι → S) [set_like.graded_monoid A] : + set_like.has_graded_smul A A := +{ smul_mem := λ i j ai bj hi hj, set_like.graded_monoid.mul_mem hi hj, } + +end subobjects + +section homogeneous_elements + +variables {S R N M : Type*} [set_like S R] [set_like N M] + +lemma set_like.is_homogeneous.graded_smul [has_add ι] [has_smul R M] {A : ι → S} {B : ι → N} + [set_like.has_graded_smul A B] {a : R} {b : M} : + set_like.is_homogeneous A a → set_like.is_homogeneous B b → set_like.is_homogeneous B (a • b) +| ⟨i, hi⟩ ⟨j, hj⟩ := ⟨i + j, set_like.has_graded_smul.smul_mem hi hj⟩ + +end homogeneous_elements diff --git a/src/algebra/module/graded_module.lean b/src/algebra/module/graded_module.lean new file mode 100644 index 0000000000000..085b9dbb85c0d --- /dev/null +++ b/src/algebra/module/graded_module.lean @@ -0,0 +1,230 @@ +/- +Copyright (c) 2022 Jujian Zhang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jujian Zhang +-/ + +import ring_theory.graded_algebra.basic +import algebra.graded_mul_action +import algebra.direct_sum.decomposition +import algebra.module.big_operators + +/-! +# Graded Module + +Given an `R`-algebra `A` graded by `𝓐`, a graded `A`-module `M` is expressed as +`direct_sum.decomposition 𝓜` and `set_like.has_graded_smul 𝓐 𝓜`. +Then `⨁ i, 𝓜 i` is an `A`-module and is isomorphic to `M`. + +## Tags + +graded module +-/ + + +section + +open_locale direct_sum + +variables {ι : Type*} (A : ι → Type*) (M : ι → Type*) + +namespace direct_sum +open graded_monoid + +/-- A graded version of `distrib_mul_action`. -/ +class gdistrib_mul_action [add_monoid ι] [gmonoid A] [Π i, add_monoid (M i)] + extends gmul_action A M := +(smul_add {i j} (a : A i) (b c : M j) : smul a (b + c) = smul a b + smul a c) +(smul_zero {i j} (a : A i) : smul a (0 : M j) = 0) + +/-- A graded version of `module`. -/ +class gmodule [add_monoid ι] [Π i, add_monoid (A i)] [Π i, add_monoid (M i)] + [gmonoid A] extends gdistrib_mul_action A M := +(add_smul {i j} (a a' : A i) (b : M j) : smul (a + a') b = smul a b + smul a' b) +(zero_smul {i j} (b : M j) : smul (0 : A i) b = 0) + +/-- A graded version of `semiring.to_module`. -/ +instance gsemiring.to_gmodule [decidable_eq ι] [add_monoid ι] + [Π (i : ι), add_comm_monoid (A i)] [gsemiring A] : + gmodule A A := +{ smul_add := λ _ _, gsemiring.mul_add, + smul_zero := λ i j, gsemiring.mul_zero, + add_smul := λ i j, gsemiring.add_mul, + zero_smul := λ i j, gsemiring.zero_mul, + ..gmonoid.to_gmul_action A } + +variables [add_monoid ι] [Π (i : ι), add_comm_monoid (A i)] [Π i, add_comm_monoid (M i)] + +/-- The piecewise multiplication from the `has_mul` instance, as a bundled homomorphism. -/ +@[simps] def gsmul_hom [gmonoid A] [gmodule A M] {i j} : + A i →+ M j →+ M (i + j) := +{ to_fun := λ a, + { to_fun := λ b, ghas_smul.smul a b, + map_zero' := gdistrib_mul_action.smul_zero _, + map_add' := gdistrib_mul_action.smul_add _ }, + map_zero' := add_monoid_hom.ext $ λ a, gmodule.zero_smul a, + map_add' := λ a₁ a₂, add_monoid_hom.ext $ λ b, gmodule.add_smul _ _ _} + +namespace gmodule + +/-- For graded monoid `A` and a graded module `M` over `A`. `gmodule.smul_add_monoid_hom` is the +`⨁ᵢ Aᵢ`-scalar multiplication on `⨁ᵢ Mᵢ` induced by `gsmul_hom`. -/ +def smul_add_monoid_hom + [decidable_eq ι] [gmonoid A] [gmodule A M] : + (⨁ i, A i) →+ (⨁ i, M i) →+ (⨁ i, M i) := +to_add_monoid $ λ i, add_monoid_hom.flip $ + to_add_monoid $ λ j, add_monoid_hom.flip $ + (of M _).comp_hom.comp $ gsmul_hom A M + +section + +open graded_monoid direct_sum gmodule + +instance [decidable_eq ι] [gmonoid A] [gmodule A M] : + has_smul (⨁ i, A i) (⨁ i, M i) := +{ smul := λ x y, smul_add_monoid_hom A M x y } + +@[simp] lemma smul_def [decidable_eq ι] [gmonoid A] [gmodule A M] + (x : ⨁ i, A i) (y : ⨁ i, M i) : x • y = smul_add_monoid_hom _ _ x y := rfl + +@[simp] lemma smul_add_monoid_hom_apply_of_of [decidable_eq ι] [gmonoid A] [gmodule A M] + {i j} (x : A i) (y : M j) : + smul_add_monoid_hom A M (direct_sum.of A i x) (of M j y) = + of M (i + j) (ghas_smul.smul x y) := +by simp [smul_add_monoid_hom] + +@[simp] lemma of_smul_of [decidable_eq ι] [gmonoid A] [gmodule A M] + {i j} (x : A i) (y : M j) : + direct_sum.of A i x • of M j y = of M (i + j) (ghas_smul.smul x y) := +smul_add_monoid_hom_apply_of_of _ _ _ _ + +open add_monoid_hom + +-- Almost identical to the proof of `direct_sum.one_mul` +private lemma one_smul [decidable_eq ι] [gmonoid A] [gmodule A M] (x : ⨁ i, M i) : + (1 : ⨁ i, A i) • x = x := +suffices smul_add_monoid_hom A M 1 = add_monoid_hom.id (⨁ i, M i), + from add_monoid_hom.congr_fun this x, +begin + apply direct_sum.add_hom_ext, intros i xi, + unfold has_one.one, + rw smul_add_monoid_hom_apply_of_of, + exact direct_sum.of_eq_of_graded_monoid_eq (one_smul (graded_monoid A) $ graded_monoid.mk i xi), +end + +-- Almost identical to the proof of `direct_sum.mul_assoc` +private lemma mul_smul [decidable_eq ι] [gsemiring A] [gmodule A M] + (a b : ⨁ i, A i) (c : ⨁ i, M i) : (a * b) • c = a • (b • c) := +suffices (smul_add_monoid_hom A M).comp_hom.comp (direct_sum.mul_hom A) + -- `λ a b c, (a * b) • c` as a bundled hom + = (add_monoid_hom.comp_hom add_monoid_hom.flip_hom $ + (smul_add_monoid_hom A M).flip.comp_hom.comp $ smul_add_monoid_hom A M).flip, + -- `λ a b c, a • (b • c)` as a bundled hom + from add_monoid_hom.congr_fun (add_monoid_hom.congr_fun (add_monoid_hom.congr_fun this a) b) c, +begin + ext ai ax bi bx ci cx : 6, + dsimp only [coe_comp, function.comp_app, comp_hom_apply_apply, flip_apply, flip_hom_apply], + rw [smul_add_monoid_hom_apply_of_of, smul_add_monoid_hom_apply_of_of, + direct_sum.mul_hom_of_of, smul_add_monoid_hom_apply_of_of], + exact direct_sum.of_eq_of_graded_monoid_eq + (mul_smul (graded_monoid.mk ai ax) (graded_monoid.mk bi bx) (graded_monoid.mk ci cx)), +end + +/-- The `module` derived from `gmodule A M`. -/ +instance module [decidable_eq ι] [gsemiring A] [gmodule A M] : + module (⨁ i, A i) (⨁ i, M i) := +{ smul := (•), + one_smul := one_smul _ _, + mul_smul := mul_smul _ _, + smul_add := λ r, (smul_add_monoid_hom A M r).map_add, + smul_zero := λ r, (smul_add_monoid_hom A M r).map_zero, + add_smul := λ r s x, by simp only [smul_def, map_add, add_monoid_hom.add_apply], + zero_smul := λ x, by simp only [smul_def, map_zero, add_monoid_hom.zero_apply] } + +end + +end gmodule + +end direct_sum + +end + +open_locale direct_sum big_operators + +variables {ι R A M σ σ' : Type*} +variables [add_monoid ι] [comm_semiring R] [semiring A] [algebra R A] +variables (𝓐 : ι → σ') [set_like σ' A] +variables (𝓜 : ι → σ) + +namespace set_like + +include σ' A σ M + +instance gmul_action [add_monoid M] [distrib_mul_action A M] + [set_like σ M] [set_like.graded_monoid 𝓐] [set_like.has_graded_smul 𝓐 𝓜] : + graded_monoid.gmul_action (λ i, 𝓐 i) (λ i, 𝓜 i) := +{ one_smul := λ ⟨i, m⟩, sigma.subtype_ext (zero_add _) (one_smul _ _), + mul_smul := λ ⟨i, a⟩ ⟨j, a'⟩ ⟨k, b⟩, sigma.subtype_ext (add_assoc _ _ _) (mul_smul _ _ _), + ..set_like.ghas_smul 𝓐 𝓜 } + +instance gdistrib_mul_action [add_monoid M] [distrib_mul_action A M] + [set_like σ M] [add_submonoid_class σ M] [set_like.graded_monoid 𝓐] + [set_like.has_graded_smul 𝓐 𝓜] : + direct_sum.gdistrib_mul_action (λ i, 𝓐 i) (λ i, 𝓜 i) := +{ smul_add := λ i j a b c, subtype.ext $ smul_add _ _ _, + smul_zero := λ i j a, subtype.ext $ smul_zero _, + ..set_like.gmul_action 𝓐 𝓜 } + +variables [add_comm_monoid M] [module A M] [set_like σ M] [add_submonoid_class σ' A] + [add_submonoid_class σ M] [set_like.graded_monoid 𝓐] [set_like.has_graded_smul 𝓐 𝓜] + +/-- `[set_like.graded_monoid 𝓐] [set_like.has_graded_smul 𝓐 𝓜]` is the internal version of graded + module, the internal version can be translated into the external version `gmodule`. -/ +instance gmodule : direct_sum.gmodule (λ i, 𝓐 i) (λ i, 𝓜 i) := +{ smul := λ i j x y, ⟨(x : A) • (y : M), set_like.has_graded_smul.smul_mem x.2 y.2⟩, + add_smul := λ i j a a' b, subtype.ext $ add_smul _ _ _, + zero_smul := λ i j b, subtype.ext $ zero_smul _ _, + ..set_like.gdistrib_mul_action 𝓐 𝓜} + +end set_like + +namespace graded_module + +include σ' A σ M + +variables [add_comm_monoid M] [module A M] [set_like σ M] [add_submonoid_class σ' A] + [add_submonoid_class σ M] [set_like.graded_monoid 𝓐] [set_like.has_graded_smul 𝓐 𝓜] + +/-- +The smul multiplication of `A` on `⨁ i, 𝓜 i` from `(⨁ i, 𝓐 i) →+ (⨁ i, 𝓜 i) →+ ⨁ i, 𝓜 i` +turns `⨁ i, 𝓜 i` into an `A`-module +-/ +def is_module [decidable_eq ι] [graded_ring 𝓐] : + module A (⨁ i, 𝓜 i) := +{ smul := λ a b, direct_sum.decompose 𝓐 a • b, + .. module.comp_hom _ (direct_sum.decompose_ring_equiv 𝓐 : A ≃+* ⨁ i, 𝓐 i).to_ring_hom } + +local attribute [instance] graded_module.is_module + +/-- +`⨁ i, 𝓜 i` and `M` are isomorphic as `A`-modules. +"The internal version" and "the external version" are isomorphism as `A`-modules. +-/ +def linear_equiv [decidable_eq ι] [graded_ring 𝓐] + [direct_sum.decomposition 𝓜] : + M ≃ₗ[A] ⨁ i, 𝓜 i := +{ to_fun := direct_sum.decompose_add_equiv 𝓜, + map_smul' := λ x y, begin + classical, + rw [← direct_sum.sum_support_decompose 𝓐 x, map_sum, finset.sum_smul, map_sum, + finset.sum_smul, finset.sum_congr rfl (λ i hi, _)], + rw [ring_hom.id_apply, ← direct_sum.sum_support_decompose 𝓜 y, map_sum, finset.smul_sum, + map_sum, finset.smul_sum, finset.sum_congr rfl (λ j hj, _)], + simp only [(•), direct_sum.decompose_add_equiv_apply, direct_sum.decompose_coe, + direct_sum.gmodule.smul_add_monoid_hom_apply_of_of], + convert direct_sum.decompose_coe 𝓜 _, + refl, + end, + .. direct_sum.decompose_add_equiv 𝓜 } + +end graded_module From bf27744463e9620ca4e4ebe951fe83530ae6949b Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Fri, 27 Jan 2023 17:10:20 +0000 Subject: [PATCH 0374/1291] feat(order/height): The height of a poset (#15026) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> Co-authored-by: Yaël Dillies Co-authored-by: Junyan Xu Co-authored-by: Junyan Xu --- src/data/list/of_fn.lean | 7 +- src/order/height.lean | 333 +++++++++++++++++++++++++++++++++++++++ 2 files changed, 339 insertions(+), 1 deletion(-) create mode 100644 src/order/height.lean diff --git a/src/data/list/of_fn.lean b/src/data/list/of_fn.lean index 4a578e41e8680..d5164175aef1f 100644 --- a/src/data/list/of_fn.lean +++ b/src/data/list/of_fn.lean @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ import data.fin.tuple.basic -import data.list.basic import data.list.join +import data.list.pairwise /-! # Lists from functions @@ -185,6 +185,11 @@ nat.rec_on n (by simp) $ λ n ihn, by simp [ihn] by simp_rw [of_fn_mul, ←of_fn_const, fin.repeat, fin.mod_nat, fin.coe_mk, add_comm, nat.add_mul_mod_self_right, nat.mod_eq_of_lt (fin.is_lt _), fin.eta] +@[simp] lemma pairwise_of_fn {R : α → α → Prop} {n} {f : fin n → α} : + (of_fn f).pairwise R ↔ ∀ ⦃i j⦄, i < j → R (f i) (f j) := +by { simp only [pairwise_iff_nth_le, fin.forall_iff, length_of_fn, nth_le_of_fn', fin.mk_lt_mk], + exact ⟨λ h i hi j hj hij, h _ _ hj hij, λ h i j hj hij, h _ (hij.trans hj) _ hj hij⟩ } + /-- Lists are equivalent to the sigma type of tuples of a given length. -/ @[simps] def equiv_sigma_tuple : list α ≃ Σ n, fin n → α := diff --git a/src/order/height.lean b/src/order/height.lean new file mode 100644 index 0000000000000..553a9744fc844 --- /dev/null +++ b/src/order/height.lean @@ -0,0 +1,333 @@ +/- +Copyright (c) 2022 Andrew Yang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Andrew Yang +-/ +import data.enat.lattice +import order.order_iso_nat +import tactic.tfae + +/-! + +# Maximal length of chains + +This file contains lemmas to work with the maximal length of strictly descending finite +sequences (chains) in a partial order. + +## Main definition + +- `set.subchain`: The set of strictly ascending lists of `α` contained in a `set α`. +- `set.chain_height`: The maximal length of a strictly ascending sequence in a partial order. +This is defined as the maximum of the lengths of `set.subchain`s, valued in `ℕ∞`. + +## Main results + +- `set.exists_chain_of_le_chain_height`: For each `n : ℕ` such that `n ≤ s.chain_height`, there + exists `s.subchain` of length `n`. +- `set.chain_height_mono`: If `s ⊆ t` then `s.chain_height ≤ t.chain_height`. +- `set.chain_height_image`: If `f` is an order embedding, then + `(f '' s).chain_height = s.chain_height`. +- `set.chain_height_insert_of_forall_lt`: If `∀ y ∈ s, y < x`, then + `(insert x s).chain_height = s.chain_height + 1`. +- `set.chain_height_insert_of_lt_forall`: If `∀ y ∈ s, x < y`, then + `(insert x s).chain_height = s.chain_height + 1`. +- `set.chain_height_union_eq`: If `∀ x ∈ s, ∀ y ∈ t, s ≤ t`, then + `(s ∪ t).chain_height = s.chain_height + t.chain_height`. +- `set.well_founded_gt_of_chain_height_ne_top`: + If `s` has finite height, then `>` is well-founded on `s`. +- `set.well_founded_lt_of_chain_height_ne_top`: + If `s` has finite height, then `<` is well-founded on `s`. + +-/ + +open list order_dual + +universes u v +variables {α β : Type*} + +namespace set +section has_lt +variables [has_lt α] [has_lt β] (s t : set α) + +/-- The set of strictly ascending lists of `α` contained in a `set α`. -/ +def subchain : set (list α) := {l | l.chain' (<) ∧ ∀ i ∈ l, i ∈ s} + +lemma nil_mem_subchain : [] ∈ s.subchain := ⟨trivial, λ x hx, hx.elim⟩ + +variables {s} {l : list α} {a : α} + +lemma cons_mem_subchain_iff : + a :: l ∈ s.subchain ↔ a ∈ s ∧ l ∈ s.subchain ∧ ∀ b ∈ l.head', a < b := +begin + refine ⟨λ h, ⟨h.2 _ (or.inl rfl), ⟨(chain'_cons'.mp h.1).2, λ i hi, h.2 _ (or.inr hi)⟩, + (chain'_cons'.mp h.1).1⟩, _⟩, + rintro ⟨h₁, h₂, h₃⟩, + split, + { rw chain'_cons', + exact ⟨h₃, h₂.1⟩ }, + { rintro i (rfl|hi), + exacts [h₁, h₂.2 _ hi] } +end + +instance : nonempty s.subchain := ⟨⟨[], s.nil_mem_subchain⟩⟩ + +variables (s) + +/-- The maximal length of a strictly ascending sequence in a partial order. -/ +noncomputable def chain_height : ℕ∞ := ⨆ l ∈ s.subchain, length l + +lemma chain_height_eq_supr_subtype : s.chain_height = ⨆ l : s.subchain, l.1.length := supr_subtype' + +lemma exists_chain_of_le_chain_height {n : ℕ} (hn : ↑n ≤ s.chain_height) : + ∃ l ∈ s.subchain, length l = n := +begin + cases (le_top : s.chain_height ≤ ⊤).eq_or_lt with ha ha; rw chain_height_eq_supr_subtype at ha, + { obtain ⟨_, ⟨⟨l, h₁, h₂⟩, rfl⟩, h₃⟩ := + not_bdd_above_iff'.mp ((with_top.supr_coe_eq_top _).mp ha) n, + exact ⟨l.take n, ⟨h₁.take _, λ x h, h₂ _ $ take_subset _ _ h⟩, + (l.length_take n).trans $ min_eq_left $ le_of_not_ge h₃⟩ }, + { rw with_top.supr_coe_lt_top at ha, + obtain ⟨⟨l, h₁, h₂⟩, e : l.length = _⟩ := nat.Sup_mem (set.range_nonempty _) ha, + refine ⟨l.take n, ⟨h₁.take _, λ x h, h₂ _ $ take_subset _ _ h⟩, + (l.length_take n).trans $ min_eq_left $ _⟩, + rwa [e, ←with_top.coe_le_coe, Sup_range, with_top.coe_supr _ ha, + ←chain_height_eq_supr_subtype] } +end + +lemma le_chain_height_tfae (n : ℕ) : + tfae [↑n ≤ s.chain_height, + ∃ l ∈ s.subchain, length l = n, + ∃ l ∈ s.subchain, n ≤ length l] := +begin + tfae_have : 1 → 2, { exact s.exists_chain_of_le_chain_height }, + tfae_have : 2 → 3, { rintro ⟨l, hls, he⟩, exact ⟨l, hls, he.ge⟩ }, + tfae_have : 3 → 1, { rintro ⟨l, hs, hn⟩, exact le_supr₂_of_le l hs (with_top.coe_le_coe.2 hn) }, + tfae_finish, +end + +variables {s t} + +lemma le_chain_height_iff {n : ℕ} : + ↑n ≤ s.chain_height ↔ ∃ l ∈ s.subchain, length l = n := +(le_chain_height_tfae s n).out 0 1 + +lemma length_le_chain_height_of_mem_subchain (hl : l ∈ s.subchain) : ↑l.length ≤ s.chain_height := +le_chain_height_iff.mpr ⟨l, hl, rfl⟩ + +lemma chain_height_eq_top_iff : s.chain_height = ⊤ ↔ ∀ n, ∃ l ∈ s.subchain, length l = n := +begin + refine ⟨λ h n, le_chain_height_iff.1 (le_top.trans_eq h.symm), λ h, _⟩, + contrapose! h, obtain ⟨n, hn⟩ := with_top.ne_top_iff_exists.1 h, + exact ⟨n + 1, λ l hs, (nat.lt_succ_iff.2 $ with_top.coe_le_coe.1 $ + (length_le_chain_height_of_mem_subchain hs).trans_eq hn.symm).ne⟩, +end + +@[simp] +lemma one_le_chain_height_iff : 1 ≤ s.chain_height ↔ s.nonempty := +begin + change ((1 : ℕ) : enat) ≤ _ ↔ _, + rw set.le_chain_height_iff, + split, + { rintro ⟨_|⟨x, xs⟩, ⟨h₁, h₂⟩, h₃⟩, + { cases h₃ }, + { exact ⟨x, h₂ _ (or.inl rfl)⟩ } }, + { rintro ⟨x, hx⟩, + exact ⟨[x], ⟨chain.nil, λ y h, (list.mem_singleton.mp h).symm ▸ hx⟩, rfl⟩ } +end + +@[simp] lemma chain_height_eq_zero_iff : s.chain_height = 0 ↔ s = ∅ := +by rw [←not_iff_not, ←ne.def, ←bot_eq_zero, ←bot_lt_iff_ne_bot, bot_eq_zero, ←enat.one_le_iff_pos, + one_le_chain_height_iff, nonempty_iff_ne_empty] + +@[simp] lemma chain_height_empty : (∅ : set α).chain_height = 0 := chain_height_eq_zero_iff.2 rfl + +@[simp] lemma chain_height_of_is_empty [is_empty α] : s.chain_height = 0 := +chain_height_eq_zero_iff.mpr (subsingleton.elim _ _) + +lemma le_chain_height_add_nat_iff {n m : ℕ} : + ↑n ≤ s.chain_height + m ↔ ∃ l ∈ s.subchain, n ≤ length l + m := +by simp_rw [← tsub_le_iff_right, ← with_top.coe_sub, (le_chain_height_tfae s (n - m)).out 0 2] + +lemma chain_height_add_le_chain_height_add (s : set α) (t : set β) (n m : ℕ) : + s.chain_height + n ≤ t.chain_height + m ↔ + ∀ l ∈ s.subchain, ∃ l' ∈ t.subchain, length l + n ≤ length l' + m := +begin + refine ⟨λ e l h, le_chain_height_add_nat_iff.1 + ((add_le_add_right (length_le_chain_height_of_mem_subchain h) _).trans e), λ H, _⟩, + by_cases s.chain_height = ⊤, + { suffices : t.chain_height = ⊤, { rw [this, top_add], exact le_top }, + rw chain_height_eq_top_iff at h ⊢, + intro k, rw (le_chain_height_tfae t k).out 1 2, + obtain ⟨l, hs, hl⟩ := h (k + m), + obtain ⟨l', ht, hl'⟩ := H l hs, + exact ⟨l', ht, (add_le_add_iff_right m).1 $ trans (hl.symm.trans_le le_self_add) hl'⟩ }, + { obtain ⟨k, hk⟩ := with_top.ne_top_iff_exists.1 h, + obtain ⟨l, hs, hl⟩ := le_chain_height_iff.1 hk.le, + rw [← hk, ← hl], + exact le_chain_height_add_nat_iff.2 (H l hs) }, +end + +lemma chain_height_le_chain_height_tfae (s : set α) (t : set β) : + tfae [s.chain_height ≤ t.chain_height, + ∀ l ∈ s.subchain, ∃ l' ∈ t.subchain, length l = length l', + ∀ l ∈ s.subchain, ∃ l' ∈ t.subchain, length l ≤ length l'] := +begin + tfae_have : 1 ↔ 3, { convert ← chain_height_add_le_chain_height_add s t 0 0; apply add_zero }, + tfae_have : 2 ↔ 3, { refine forall₂_congr (λ l hl, _), + simp_rw [← (le_chain_height_tfae t l.length).out 1 2, eq_comm] }, + tfae_finish +end + +lemma chain_height_le_chain_height_iff {t : set β} : + s.chain_height ≤ t.chain_height ↔ + ∀ l ∈ s.subchain, ∃ l' ∈ t.subchain, length l = length l' := +(chain_height_le_chain_height_tfae s t).out 0 1 + +lemma chain_height_le_chain_height_iff_le {t : set β} : + s.chain_height ≤ t.chain_height ↔ + ∀ l ∈ s.subchain, ∃ l' ∈ t.subchain, length l ≤ length l' := +(chain_height_le_chain_height_tfae s t).out 0 2 + +lemma chain_height_mono (h : s ⊆ t) : s.chain_height ≤ t.chain_height := +chain_height_le_chain_height_iff.2 $ λ l hl, ⟨l, ⟨hl.1, λ i hi, h $ hl.2 i hi⟩, rfl⟩ + +lemma chain_height_image + (f : α → β) (hf : ∀ {x y}, x < y ↔ f x < f y) (s : set α) : + (f '' s).chain_height = s.chain_height := +begin + apply le_antisymm; rw chain_height_le_chain_height_iff, + { suffices : ∀ l ∈ (f '' s).subchain, ∃ l' ∈ s.subchain, map f l' = l, + { intros l hl, obtain ⟨l', h₁, rfl⟩ := this l hl, exact ⟨l', h₁, length_map _ _⟩ }, + intro l, + induction l with x xs hx, + { exact λ _, ⟨nil, ⟨trivial, λ _ h, h.elim⟩, rfl⟩ }, + { intros h, + rw cons_mem_subchain_iff at h, + obtain ⟨⟨x, hx', rfl⟩, h₁, h₂⟩ := h, + obtain ⟨l', h₃, rfl⟩ := hx h₁, + refine ⟨x :: l', set.cons_mem_subchain_iff.mpr ⟨hx', h₃, _⟩, rfl⟩, + cases l', { simp }, { simpa [← hf] using h₂ } } }, + { intros l hl, + refine ⟨l.map f, ⟨_, _⟩, _⟩, + { simp_rw [chain'_map, ← hf], exact hl.1 }, + { intros _ e, obtain ⟨a, ha, rfl⟩ := mem_map.mp e, exact set.mem_image_of_mem _ (hl.2 _ ha) }, + { rw length_map } }, +end + +variables (s) + +@[simp] lemma chain_height_dual : (of_dual ⁻¹' s).chain_height = s.chain_height := +begin + apply le_antisymm; + { rw chain_height_le_chain_height_iff, + rintro l ⟨h₁, h₂⟩, + exact ⟨l.reverse, ⟨chain'_reverse.mpr h₁, + λ i h, h₂ i (mem_reverse.mp h)⟩, (length_reverse _).symm⟩ } +end + +end has_lt + +section preorder +variables (s t : set α) [preorder α] + +lemma chain_height_eq_supr_Ici : s.chain_height = ⨆ i ∈ s, (s ∩ set.Ici i).chain_height := +begin + apply le_antisymm, + { refine supr₂_le _, + rintro (_ | ⟨x, xs⟩) h, + { exact zero_le _ }, + { apply le_trans _ (le_supr₂ x (cons_mem_subchain_iff.mp h).1), + apply length_le_chain_height_of_mem_subchain, + refine ⟨h.1, λ i hi, ⟨h.2 i hi, _⟩⟩, + cases hi, { exact hi.symm.le }, + cases chain'_iff_pairwise.mp h.1 with _ _ h', + exact (h' _ hi).le } }, + { exact supr₂_le (λ i hi, chain_height_mono $ set.inter_subset_left _ _) } +end + +lemma chain_height_eq_supr_Iic : s.chain_height = ⨆ i ∈ s, (s ∩ set.Iic i).chain_height := +by { simp_rw ←chain_height_dual (_ ∩ _), rw [←chain_height_dual, chain_height_eq_supr_Ici], refl } + +variables {s t} + +lemma chain_height_insert_of_forall_gt (a : α) (hx : ∀ b ∈ s, a < b) : + (insert a s).chain_height = s.chain_height + 1 := +begin + rw ← add_zero (insert a s).chain_height, + change (insert a s).chain_height + (0 : ℕ) = s.chain_height + (1 : ℕ), + apply le_antisymm; rw chain_height_add_le_chain_height_add, + { rintro (_|⟨y, ys⟩) h, + { exact ⟨[], nil_mem_subchain _, zero_le _⟩ }, + { have h' := cons_mem_subchain_iff.mp h, + refine ⟨ys, ⟨h'.2.1.1, λ i hi, _⟩, by simp⟩, + apply (h'.2.1.2 i hi).resolve_left, + rintro rfl, + cases chain'_iff_pairwise.mp h.1 with _ _ hy, + cases h'.1 with h' h', + exacts [(hy _ hi).ne h', not_le_of_gt (hy _ hi) (hx _ h').le] } }, + { intros l hl, + refine ⟨a :: l, ⟨_, _⟩, by simp⟩, + { rw chain'_cons', exact ⟨λ y hy, hx _ (hl.2 _ (mem_of_mem_head' hy)), hl.1⟩ }, + { rintro x (rfl|hx), exacts [or.inl (set.mem_singleton x), or.inr (hl.2 x hx)] } } +end + +lemma chain_height_insert_of_forall_lt (a : α) (ha : ∀ b ∈ s, b < a) : + (insert a s).chain_height = s.chain_height + 1 := +by { rw [←chain_height_dual, ←chain_height_dual s], exact chain_height_insert_of_forall_gt _ ha } + +lemma chain_height_union_le : (s ∪ t).chain_height ≤ s.chain_height + t.chain_height := +begin + classical, + refine supr₂_le (λ l hl, _), + let l₁ := l.filter (∈ s), let l₂ := l.filter (∈ t), + have hl₁ : ↑l₁.length ≤ s.chain_height, + { apply set.length_le_chain_height_of_mem_subchain, + exact ⟨hl.1.sublist (filter_sublist _), λ i h, (of_mem_filter h : _)⟩ }, + have hl₂ : ↑l₂.length ≤ t.chain_height, + { apply set.length_le_chain_height_of_mem_subchain, + exact ⟨hl.1.sublist (filter_sublist _), λ i h, (of_mem_filter h : _)⟩ }, + refine le_trans _ (add_le_add hl₁ hl₂), + simp_rw [← with_top.coe_add, with_top.coe_le_coe, ← multiset.coe_card, + ← multiset.card_add, ← multiset.coe_filter], + rw [multiset.filter_add_filter, multiset.filter_eq_self.mpr, multiset.card_add], + exacts [le_add_right rfl.le, hl.2] +end + +lemma chain_height_union_eq (s t : set α) (H : ∀ (a ∈ s) (b ∈ t), a < b) : + (s ∪ t).chain_height = s.chain_height + t.chain_height := +begin + cases h : t.chain_height, + { rw [with_top.none_eq_top, add_top, eq_top_iff, ← with_top.none_eq_top, ← h], + exact set.chain_height_mono (set.subset_union_right _ _) }, + apply le_antisymm, + { rw ← h, + exact chain_height_union_le }, + rw [with_top.some_eq_coe, ← add_zero (s ∪ t).chain_height, ← with_top.coe_zero, + chain_height_add_le_chain_height_add], + intros l hl, + obtain ⟨l', hl', rfl⟩ := exists_chain_of_le_chain_height t h.symm.le, + refine ⟨l ++ l', ⟨chain'.append hl.1 hl'.1 $ λ x hx y hy, _, λ i hi, _⟩, by simp⟩, + { exact H x (hl.2 _ $ mem_of_mem_last' hx) y (hl'.2 _ $ mem_of_mem_head' hy) }, + { rw mem_append at hi, cases hi, exacts [or.inl (hl.2 _ hi), or.inr (hl'.2 _ hi)] } +end + +lemma well_founded_gt_of_chain_height_ne_top (s : set α) (hs : s.chain_height ≠ ⊤) : + well_founded_gt s := +begin + obtain ⟨n, hn⟩ := with_top.ne_top_iff_exists.1 hs, + refine ⟨rel_embedding.well_founded_iff_no_descending_seq.2 ⟨λ f, _⟩⟩, + refine n.lt_succ_self.not_le (with_top.coe_le_coe.1 $ hn.symm ▸ _), + refine le_supr₂_of_le _ ⟨chain'_map_of_chain' coe (λ _ _, id) + (chain'_iff_pairwise.2 $ pairwise_of_fn.2 $ λ i j, f.map_rel_iff.2), λ i h, _⟩ _, + { exact n.succ }, + { obtain ⟨a, ha, rfl⟩ := mem_map.1 h, exact a.prop }, + { rw [length_map, length_of_fn], exact le_rfl }, +end + +lemma well_founded_lt_of_chain_height_ne_top (s : set α) (hs : s.chain_height ≠ ⊤) : + well_founded_lt s := +well_founded_gt_of_chain_height_ne_top (of_dual ⁻¹' s) $ by rwa chain_height_dual + +end preorder + +end set From b86c528d08a52a1fdb50d999232408e1c7e85d7d Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Fri, 27 Jan 2023 17:10:22 +0000 Subject: [PATCH 0375/1291] feat(ring_theory/ideal/local_ring): API for local rings. (#17185) Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> Co-authored-by: Junyan Xu Co-authored-by: Junyan Xu Co-authored-by: Eric Wieser --- .../prime_spectrum/basic.lean | 30 +++++++++++++++- src/ring_theory/discrete_valuation_ring.lean | 2 +- src/ring_theory/ideal/basic.lean | 35 +++++++++++++------ src/ring_theory/ideal/local_ring.lean | 27 ++++++++++++++ src/ring_theory/ideal/operations.lean | 2 +- src/ring_theory/localization/at_prime.lean | 7 ++-- 6 files changed, 88 insertions(+), 15 deletions(-) diff --git a/src/algebraic_geometry/prime_spectrum/basic.lean b/src/algebraic_geometry/prime_spectrum/basic.lean index ba43a6c38d108..711ffcd6a0da6 100644 --- a/src/algebraic_geometry/prime_spectrum/basic.lean +++ b/src/algebraic_geometry/prime_spectrum/basic.lean @@ -446,7 +446,7 @@ begin (is_closed_singleton_iff_is_maximal _).1 (t1_space.t1 ⟨⊥, hbot⟩)) (not_not.2 rfl)) }, { refine ⟨λ x, (is_closed_singleton_iff_is_maximal x).2 _⟩, by_cases hx : x.as_ideal = ⊥, - { exact hx.symm ▸ @ideal.bot_is_maximal R (@field.to_division_ring _ h.to_field) }, + { letI := h.to_field, exact hx.symm ▸ ideal.bot_is_maximal }, { exact absurd h (ring.not_is_field_iff_exists_prime.2 ⟨x.as_ideal, ⟨hx, x.2⟩⟩) } } end @@ -822,6 +822,14 @@ order_embedding.of_map_le_iff nhds $ λ a b, (le_iff_specializes a b).symm instance : t0_space (prime_spectrum R) := ⟨nhds_order_embedding.injective⟩ +instance [is_domain R] : order_bot (prime_spectrum R) := +{ bot := ⟨⊥, ideal.bot_prime⟩, + bot_le := λ I, @bot_le _ _ _ I.as_ideal } + +instance {R : Type*} [field R] : unique (prime_spectrum R) := +{ default := ⊥, + uniq := λ x, ext _ _ ((is_simple_order.eq_bot_or_eq_top _).resolve_right x.2.ne_top) } + end order /-- If `x` specializes to `y`, then there is a natural map from the localization of `y` to the @@ -856,4 +864,24 @@ by { rw [(local_hom_tfae f).out 0 4, prime_spectrum.ext_iff], refl } [is_local_ring_hom f] : prime_spectrum.comap f (closed_point S) = closed_point R := (is_local_ring_hom_iff_comap_closed_point f).mp infer_instance +lemma specializes_closed_point (x : prime_spectrum R) : + x ⤳ closed_point R := +(prime_spectrum.le_iff_specializes _ _).mp (local_ring.le_maximal_ideal x.2.1) + +lemma closed_point_mem_iff (U : topological_space.opens $ prime_spectrum R) : + closed_point R ∈ U ↔ U = ⊤ := +begin + split, + { rw eq_top_iff, exact λ h x _, (specializes_closed_point x).mem_open U.2 h }, + { rintro rfl, trivial } +end + +@[simp] lemma _root_.prime_spectrum.comap_residue (x : prime_spectrum (residue_field R)) : + prime_spectrum.comap (residue R) x = closed_point R := +begin + rw subsingleton.elim x ⊥, + ext1, + exact ideal.mk_ker, +end + end local_ring diff --git a/src/ring_theory/discrete_valuation_ring.lean b/src/ring_theory/discrete_valuation_ring.lean index 9450069c0d4b8..5c2df165d347a 100644 --- a/src/ring_theory/discrete_valuation_ring.lean +++ b/src/ring_theory/discrete_valuation_ring.lean @@ -61,7 +61,7 @@ lemma not_a_field : maximal_ideal R ≠ ⊥ := not_a_field' /-- A discrete valuation ring `R` is not a field. -/ lemma not_is_field : ¬ is_field R := -ring.not_is_field_iff_exists_prime.mpr ⟨_, not_a_field R, is_maximal.is_prime' (maximal_ideal R)⟩ +local_ring.is_field_iff_maximal_ideal_eq.not.mpr (not_a_field R) variable {R} diff --git a/src/ring_theory/ideal/basic.lean b/src/ring_theory/ideal/basic.lean index fd1ffa938c20c..9a123aa170a2f 100644 --- a/src/ring_theory/ideal/basic.lean +++ b/src/ring_theory/ideal/basic.lean @@ -535,12 +535,12 @@ end ideal end ring -section division_ring -variables {K : Type u} [division_ring K] (I : ideal K) +section division_semiring +variables {K : Type u} [division_semiring K] (I : ideal K) namespace ideal -/-- All ideals in a division ring are trivial. -/ +/-- All ideals in a division (semi)ring are trivial. -/ lemma eq_bot_or_top : I = ⊥ ∨ I = ⊤ := begin rw or_iff_not_imp_right, @@ -553,8 +553,8 @@ begin simpa [H, h1] using I.mul_mem_left r⁻¹ hr, end -/-- Ideals of a `division_ring` are a simple order. Thanks to the way abbreviations work, this -automatically gives a `is_simple_module K` instance. -/ +/-- Ideals of a `division_semiring` are a simple order. Thanks to the way abbreviations work, +this automatically gives a `is_simple_module K` instance. -/ instance : is_simple_order (ideal K) := ⟨eq_bot_or_top⟩ lemma eq_bot_of_prime [h : I.is_prime] : I = ⊥ := @@ -566,7 +566,7 @@ lemma bot_is_maximal : is_maximal (⊥ : ideal K) := end ideal -end division_ring +end division_semiring section comm_ring @@ -583,12 +583,11 @@ end ideal end comm_ring +-- TODO: consider moving the lemmas below out of the `ring` namespace since they are +-- about `comm_semiring`s. namespace ring -variables {R : Type*} [comm_ring R] - -lemma not_is_field_of_subsingleton {R : Type*} [ring R] [subsingleton R] : ¬ is_field R := -λ ⟨⟨x, y, hxy⟩, _, _⟩, hxy (subsingleton.elim x y) +variables {R : Type*} [comm_semiring R] lemma exists_not_is_unit_of_not_is_field [nontrivial R] (hf : ¬ is_field R) : ∃ x ≠ (0 : R), ¬ is_unit x := @@ -624,6 +623,22 @@ not_is_field_iff_exists_ideal_bot_lt_and_lt_top.trans ⟨p, bot_lt_iff_ne_bot.mp (lt_of_lt_of_le bot_lt le_p), hp.is_prime⟩, λ ⟨p, ne_bot, prime⟩, ⟨p, bot_lt_iff_ne_bot.mpr ne_bot, lt_top_iff_ne_top.mpr prime.1⟩⟩ +/-- Also see `ideal.is_simple_order` for the forward direction as an instance when `R` is a +division (semi)ring. + +This result actually holds for all division semirings, but we lack the predicate to state it. -/ +lemma is_field_iff_is_simple_order_ideal : + is_field R ↔ is_simple_order (ideal R) := +begin + casesI subsingleton_or_nontrivial R, + { exact ⟨λ h, (not_is_field_of_subsingleton _ h).elim, + λ h, by exactI (false_of_nontrivial_of_subsingleton $ ideal R).elim⟩ }, + rw [← not_iff_not, ring.not_is_field_iff_exists_ideal_bot_lt_and_lt_top, ← not_iff_not], + push_neg, + simp_rw [lt_top_iff_ne_top, bot_lt_iff_ne_bot, ← or_iff_not_imp_left, not_ne_iff], + exact ⟨λ h, ⟨h⟩, λ h, h.2⟩ +end + /-- When a ring is not a field, the maximal ideals are nontrivial. -/ lemma ne_bot_of_is_maximal_of_not_is_field [nontrivial R] {M : ideal R} (max : M.is_maximal) (not_field : ¬ is_field R) : M ≠ ⊥ := diff --git a/src/ring_theory/ideal/local_ring.lean b/src/ring_theory/ideal/local_ring.lean index 258f38590b2e2..c3d77a19bdc4c 100644 --- a/src/ring_theory/ideal/local_ring.lean +++ b/src/ring_theory/ideal/local_ring.lean @@ -130,6 +130,11 @@ end @[simp] lemma mem_maximal_ideal (x) : x ∈ maximal_ideal R ↔ x ∈ nonunits R := iff.rfl +lemma is_field_iff_maximal_ideal_eq : + is_field R ↔ maximal_ideal R = ⊥ := +not_iff_not.mp ⟨ring.ne_bot_of_is_maximal_of_not_is_field infer_instance, + λ h, ring.not_is_field_iff_exists_prime.mpr ⟨_, h, ideal.is_maximal.is_prime' _⟩⟩ + end local_ring end comm_semiring @@ -324,10 +329,28 @@ ideal.quotient.algebra _ lemma residue_field.algebra_map_eq : algebra_map R (residue_field R) = residue R := rfl +instance : is_local_ring_hom (local_ring.residue R) := +⟨λ a ha, not_not.mp (ideal.quotient.eq_zero_iff_mem.not.mp (is_unit_iff_ne_zero.mp ha))⟩ + variables {R} namespace residue_field +/-- A local ring homomorphism into a field can be descended onto the residue field. -/ +def lift {R S : Type*} [comm_ring R] [local_ring R] [field S] + (f : R →+* S) [is_local_ring_hom f] : local_ring.residue_field R →+* S := +ideal.quotient.lift _ f (λ a ha, + classical.by_contradiction (λ h, ha (is_unit_of_map_unit f a (is_unit_iff_ne_zero.mpr h)))) + +lemma lift_comp_residue {R S : Type*} [comm_ring R] [local_ring R] [field S] (f : R →+* S) + [is_local_ring_hom f] : (lift f).comp (residue R) = f := +ring_hom.ext (λ _, rfl) + +@[simp] +lemma lift_residue_apply {R S : Type*} [comm_ring R] [local_ring R] [field S] (f : R →+* S) + [is_local_ring_hom f] (x) : lift f (residue R x) = f x := +rfl + /-- The map on residue fields induced by a local homomorphism between local rings -/ def map (f : R →+* S) [is_local_ring_hom f] : residue_field R →+* residue_field S := ideal.quotient.lift (maximal_ideal R) ((ideal.quotient.mk _).comp f) $ @@ -433,3 +456,7 @@ local_ring.of_is_unit_or_is_unit_one_sub_self $ λ a, else or.inl $ is_unit.mk0 a h end field + +lemma local_ring.maximal_ideal_eq_bot {R : Type*} [field R] : + local_ring.maximal_ideal R = ⊥ := +local_ring.is_field_iff_maximal_ideal_eq.mp (field.to_is_field R) diff --git a/src/ring_theory/ideal/operations.lean b/src/ring_theory/ideal/operations.lean index 1d2ea816a284b..cbb6eebef923e 100644 --- a/src/ring_theory/ideal/operations.lean +++ b/src/ring_theory/ideal/operations.lean @@ -1904,7 +1904,7 @@ end (⊥ : ideal (R ⧸ I)).is_maximal ↔ I.is_maximal := ⟨λ hI, (@mk_ker _ _ I) ▸ @comap_is_maximal_of_surjective _ _ _ _ _ _ (quotient.mk I) quotient.mk_surjective ⊥ hI, - λ hI, @bot_is_maximal _ (@field.to_division_ring _ (@quotient.field _ _ I hI)) ⟩ + λ hI, by { resetI, letI := quotient.field I, exact bot_is_maximal }⟩ /-- See also `ideal.mem_quotient_iff_mem` in case `I ≤ J`. -/ @[simp] diff --git a/src/ring_theory/localization/at_prime.lean b/src/ring_theory/localization/at_prime.lean index f37c5f099febf..1b3a5bc1a1418 100644 --- a/src/ring_theory/localization/at_prime.lean +++ b/src/ring_theory/localization/at_prime.lean @@ -139,6 +139,10 @@ not_iff_not.mp $ by simpa only [local_ring.mem_maximal_ideal, mem_nonunits_iff, not_not] using is_unit_to_map_iff S I x +lemma comap_maximal_ideal (h : _root_.local_ring S := local_ring S I) : + (local_ring.maximal_ideal S).comap (algebra_map R S) = I := +ideal.ext $ λ x, by simpa only [ideal.mem_comap] using to_map_mem_maximal_iff _ I x + lemma is_unit_mk'_iff (x : R) (y : I.prime_compl) : is_unit (mk' S x y) ↔ x ∈ I.prime_compl := ⟨λ h hx, mk'_mem_iff.mpr ((to_map_mem_maximal_iff S I x).mpr hx) h, @@ -168,8 +172,7 @@ variables {I} lemma at_prime.comap_maximal_ideal : ideal.comap (algebra_map R (localization.at_prime I)) (local_ring.maximal_ideal (localization I.prime_compl)) = I := -ideal.ext $ λ x, by -simpa only [ideal.mem_comap] using at_prime.to_map_mem_maximal_iff _ I x +at_prime.comap_maximal_ideal _ _ /-- The image of `I` in the localization at `I.prime_compl` is a maximal ideal, and in particular it is the unique maximal ideal given by the local ring structure `at_prime.local_ring` -/ From 97f079b7e89566de3a1143f887713667328c38ba Mon Sep 17 00:00:00 2001 From: JovanGerb Date: Fri, 27 Jan 2023 17:10:24 +0000 Subject: [PATCH 0376/1291] chore(topology/metric_space/anti_lipschitz): fix typo (#18263) Fix a typo in the definition of `antilipschitz_with`. I presume that the definition is correct, and was not intended as described in the docstring. --- src/topology/metric_space/antilipschitz.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/topology/metric_space/antilipschitz.lean b/src/topology/metric_space/antilipschitz.lean index f07c0959775e0..1942bce145101 100644 --- a/src/topology/metric_space/antilipschitz.lean +++ b/src/topology/metric_space/antilipschitz.lean @@ -26,7 +26,7 @@ open_locale nnreal ennreal uniformity open set filter bornology /-- We say that `f : α → β` is `antilipschitz_with K` if for any two points `x`, `y` we have -`K * edist x y ≤ edist (f x) (f y)`. -/ +`edist x y ≤ K * edist (f x) (f y)`. -/ def antilipschitz_with [pseudo_emetric_space α] [pseudo_emetric_space β] (K : ℝ≥0) (f : α → β) := ∀ x y, edist x y ≤ K * edist (f x) (f y) From 5fe298160aa02b0f3cf95690a1265232cdd9563c Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Fri, 27 Jan 2023 17:10:25 +0000 Subject: [PATCH 0377/1291] feat(ring_theory): lemmas on algebra norm (#18298) This PR includes some results on the algebra norm needed for the ideal norm: * `algebra.norm R (0 : S) = 0` under weaker conditions on `R` and `S` * use `module.free` and `module.finite` instead of explicit bases in `algebra.norm_eq_zero_iff` / generalize from vector spaces over a field to free modules over a ring * the norm map between localizations is equal to the norm over the base fields --- src/linear_algebra/free_module/basic.lean | 5 +- src/ring_theory/localization/norm.lean | 55 +++++++++++++++++++++ src/ring_theory/norm.lean | 58 ++++++++++++++--------- 3 files changed, 95 insertions(+), 23 deletions(-) create mode 100644 src/ring_theory/localization/norm.lean diff --git a/src/linear_algebra/free_module/basic.lean b/src/linear_algebra/free_module/basic.lean index 25c1a2c1321fe..aa89253c1406a 100644 --- a/src/linear_algebra/free_module/basic.lean +++ b/src/linear_algebra/free_module/basic.lean @@ -68,7 +68,7 @@ variables [add_comm_monoid N] [module R N] /-- If `module.free R M` then `choose_basis_index R M` is the `ι` which indexes the basis `ι → M`. -/ -@[nolint has_nonempty_instance] def choose_basis_index := (exists_basis R M).some.1 +def choose_basis_index := (exists_basis R M).some.1 /-- If `module.free R M` then `choose_basis : ι → M` is the basis. Here `ι = choose_basis_index R M`. -/ @@ -92,6 +92,9 @@ noncomputable def constr {S : Type z} [semiring S] [module S N] [smul_comm_class instance no_zero_smul_divisors [no_zero_divisors R] : no_zero_smul_divisors R M := let ⟨⟨_, b⟩⟩ := exists_basis R M in b.no_zero_smul_divisors +instance [nontrivial M] : nonempty (module.free.choose_basis_index R M) := +(module.free.choose_basis R M).index_nonempty + variables {R M N} lemma of_equiv (e : M ≃ₗ[R] N) : module.free R N := diff --git a/src/ring_theory/localization/norm.lean b/src/ring_theory/localization/norm.lean new file mode 100644 index 0000000000000..0e3b7a7d28a36 --- /dev/null +++ b/src/ring_theory/localization/norm.lean @@ -0,0 +1,55 @@ +/- +Copyright (c) 2023 Anne Baanen. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Anne Baanen +-/ + +import ring_theory.localization.module +import ring_theory.norm + +/-! + +# Field/algebra norm and localization + +This file contains results on the combination of `algebra.norm` and `is_localization`. + +## Main results + + * `algebra.norm_localization`: let `S` be an extension of `R` and `Rₘ Sₘ` be localizations at `M` + of `R S` respectively. Then the norm of `a : Sₘ` over `Rₘ` is the norm of `a : S` over `R` + if `S` is free as `R`-module + +## Tags + +field norm, algebra norm, localization + +-/ + +open_locale non_zero_divisors + +variables (R : Type*) {S : Type*} [comm_ring R] [comm_ring S] [algebra R S] +variables {Rₘ Sₘ : Type*} [comm_ring Rₘ] [algebra R Rₘ] [comm_ring Sₘ] [algebra S Sₘ] +variables {M : submonoid R} +variables [is_localization M Rₘ] [is_localization (algebra.algebra_map_submonoid S M) Sₘ] +variables [algebra Rₘ Sₘ] [algebra R Sₘ] [is_scalar_tower R Rₘ Sₘ] [is_scalar_tower R S Sₘ] + +/-- Let `S` be an extension of `R` and `Rₘ Sₘ` be localizations at `M` of `R S` respectively. +Then the norm of `a : Sₘ` over `Rₘ` is the norm of `a : S` over `R` if `S` is free as `R`-module. +-/ +lemma algebra.norm_localization [module.free R S] [module.finite R S] + (a : S) (hM : algebra.algebra_map_submonoid S M ≤ S⁰) : + algebra.norm Rₘ (algebra_map S Sₘ a) = algebra_map R Rₘ (algebra.norm R a) := +begin + casesI subsingleton_or_nontrivial R, + { haveI : subsingleton Rₘ := module.subsingleton R Rₘ, + simp }, + let b := module.free.choose_basis R S, + letI := classical.dec_eq (module.free.choose_basis_index R S), + rw [algebra.norm_eq_matrix_det (b.localization_localization Rₘ M Sₘ hM), + algebra.norm_eq_matrix_det b, ring_hom.map_det], + congr, + ext i j, + simp only [matrix.map_apply, ring_hom.map_matrix_apply, algebra.left_mul_matrix_eq_repr_mul, + basis.localization_localization_apply, ← _root_.map_mul], + apply basis.localization_localization_repr_algebra_map +end diff --git a/src/ring_theory/norm.lean b/src/ring_theory/norm.lean index ec9bca31b5aac..8c1c231ea7500 100644 --- a/src/ring_theory/norm.lean +++ b/src/ring_theory/norm.lean @@ -130,42 +130,56 @@ end eq_prod_roots section eq_zero_iff variables [finite ι] -lemma norm_eq_zero_iff_of_basis [is_domain R] [is_domain S] (b : basis ι R S) {x : S} : - algebra.norm R x = 0 ↔ x = 0 := +@[simp] lemma norm_zero [nontrivial S] [module.free R S] [module.finite R S] : + norm R (0 : S) = 0 := +begin + nontriviality, + rw [norm_apply, coe_lmul_eq_mul, map_zero, linear_map.det_zero' (module.free.choose_basis R S)] +end + +@[simp] lemma norm_eq_zero_iff [is_domain R] [is_domain S] [module.free R S] [module.finite R S] + {x : S} : + norm R x = 0 ↔ x = 0 := begin - casesI nonempty_fintype ι, - have hι : nonempty ι := b.index_nonempty, - letI := classical.dec_eq ι, - rw algebra.norm_eq_matrix_det b, split, - { rw ← matrix.exists_mul_vec_eq_zero_iff, + let b := module.free.choose_basis R S, + swap, { rintro rfl, exact norm_zero }, + { letI := classical.dec_eq (module.free.choose_basis_index R S), + rw [norm_eq_matrix_det b, + ← matrix.exists_mul_vec_eq_zero_iff], rintros ⟨v, v_ne, hv⟩, rw [← b.equiv_fun.apply_symm_apply v, b.equiv_fun_symm_apply, b.equiv_fun_apply, - algebra.left_mul_matrix_mul_vec_repr] at hv, + left_mul_matrix_mul_vec_repr] at hv, refine (mul_eq_zero.mp (b.ext_elem $ λ i, _)).resolve_right (show ∑ i, v i • b i ≠ 0, from _), { simpa only [linear_equiv.map_zero, pi.zero_apply] using congr_fun hv i }, { contrapose! v_ne with sum_eq, apply b.equiv_fun.symm.injective, rw [b.equiv_fun_symm_apply, sum_eq, linear_equiv.map_zero] } }, - { rintro rfl, - rw [alg_hom.map_zero, matrix.det_zero hι] }, end -lemma norm_ne_zero_iff_of_basis [is_domain R] [is_domain S] (b : basis ι R S) {x : S} : - algebra.norm R x ≠ 0 ↔ x ≠ 0 := -not_iff_not.mpr (algebra.norm_eq_zero_iff_of_basis b) - -/-- See also `algebra.norm_eq_zero_iff'` if you already have rewritten with `algebra.norm_apply`. -/ -@[simp] -lemma norm_eq_zero_iff {K L : Type*} [field K] [comm_ring L] [algebra K L] [is_domain L] - [finite_dimensional K L] {x : L} : algebra.norm K x = 0 ↔ x = 0 := -algebra.norm_eq_zero_iff_of_basis (basis.of_vector_space K L) +lemma norm_ne_zero_iff [is_domain R] [is_domain S] [module.free R S] [module.finite R S] + {x : S} : + norm R x ≠ 0 ↔ x ≠ 0 := +not_iff_not.mpr norm_eq_zero_iff /-- This is `algebra.norm_eq_zero_iff` composed with `algebra.norm_apply`. -/ @[simp] -lemma norm_eq_zero_iff' {K L : Type*} [field K] [comm_ring L] [algebra K L] [is_domain L] - [finite_dimensional K L] {x : L} : linear_map.det (linear_map.mul K L x) = 0 ↔ x = 0 := -algebra.norm_eq_zero_iff_of_basis (basis.of_vector_space K L) +lemma norm_eq_zero_iff' [is_domain R] [is_domain S] [module.free R S] [module.finite R S] + {x : S} : + linear_map.det (linear_map.mul R S x) = 0 ↔ x = 0 := +norm_eq_zero_iff + +lemma norm_eq_zero_iff_of_basis [is_domain R] [is_domain S] (b : basis ι R S) {x : S} : + algebra.norm R x = 0 ↔ x = 0 := +begin + haveI : module.free R S := module.free.of_basis b, + haveI : module.finite R S := module.finite.of_basis b, + exact norm_eq_zero_iff +end + +lemma norm_ne_zero_iff_of_basis [is_domain R] [is_domain S] (b : basis ι R S) {x : S} : + algebra.norm R x ≠ 0 ↔ x ≠ 0 := +not_iff_not.mpr (norm_eq_zero_iff_of_basis b) end eq_zero_iff From feb99064803fd3108e37c18b0f77d0a8344677a3 Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Fri, 27 Jan 2023 23:12:00 +0000 Subject: [PATCH 0378/1291] fix(data/set_like): unbundled subclasses of `out_param` classes should not repeat the parents' `out_param` (#18291) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR is a backport for a mathlib4 fix for a large amount of issues encountered in the port of `group_theory.subgroup.basic`, leanprover-community/mathlib4#1797. Subobject classes such as `mul_mem_class` and `submonoid_class` take a `set_like` parameter, and we were used to just copy over the `M : out_param Type` declaration from `set_like`. Since Lean 3 synthesizes from left to right, we'd fill in the `out_param` from `set_like`, then the `has_mul M` instance would be available for synthesis and we'd be in business. In Lean 4 however, we will often go from right to left, so that `MulMemClass ?M S [?inst : Mul ?M]` is handled first, which can't be solved before finding a `Mul ?M` instance on the metavariable `?M`, causing search through all `Mul` instances. The solution is: whenever a class is declared that takes another class as parameter (i.e. unbundled inheritance), the `out_param`s of the parent class should be unmarked and become in-params in the child class. Then Lean will try to find the parent class instance first, fill in the `out_param`s and we'll be able to synthesize the child class instance without problems. One consequence is that sometimes we have to give slightly more type ascription when talking about membership and the types don't quite align: if `M` and `M'` are semireducibly defeq, then before `zero_mem_class S M` would work to prove `(0 : M') ∈ (s : S)`, since the `out_param` became a metavariable, was filled in, and then checked (up to semireducibility apparently) for equality. Now `M` is checked to equal `M'` before applying the instance, with instance-reducible transparency. I don't think this is a huge issue since it feels Lean 4 is stricter about these kinds of equalities anyway. Mathlib4 pair: leanprover-community/mathlib4#1832 --- src/algebra/module/submodule/basic.lean | 7 +++++-- src/algebra/monoid_algebra/grading.lean | 7 ++++--- src/data/set_like/basic.lean | 9 +++++++++ src/field_theory/subfield.lean | 2 +- .../group_action/sub_mul_action.lean | 18 ++++++++++++++---- src/group_theory/submonoid/basic.lean | 8 ++++---- src/group_theory/submonoid/pointwise.lean | 2 +- src/group_theory/subsemigroup/basic.lean | 4 ++-- src/logic/basic.lean | 4 ++-- src/ring_theory/derivation.lean | 4 ++-- .../non_unital_subsemiring/basic.lean | 4 ++-- src/ring_theory/subring/basic.lean | 4 ++-- src/ring_theory/subsemiring/basic.lean | 6 +++--- 13 files changed, 51 insertions(+), 28 deletions(-) diff --git a/src/algebra/module/submodule/basic.lean b/src/algebra/module/submodule/basic.lean index c476f739d5f20..18cc10c629849 100644 --- a/src/algebra/module/submodule/basic.lean +++ b/src/algebra/module/submodule/basic.lean @@ -32,8 +32,11 @@ variables {G : Type u''} {S : Type u'} {R : Type u} {M : Type v} {ι : Type w} set_option old_structure_cmd true -/-- `submodule_class S R M` says `S` is a type of submodules `s ≤ M`. -/ -class submodule_class (S : Type*) (R M : out_param $ Type*) [add_zero_class M] +/-- `submodule_class S R M` says `S` is a type of submodules `s ≤ M`. + +Note that only `R` is marked as `out_param` since `M` is already supplied by the `set_like` class. +-/ +class submodule_class (S : Type*) (R : out_param $ Type*) (M : Type*) [add_zero_class M] [has_smul R M] [set_like S M] [add_submonoid_class S M] extends smul_mem_class S R M /-- A submodule of a module is one which is closed under vector operations. diff --git a/src/algebra/monoid_algebra/grading.lean b/src/algebra/monoid_algebra/grading.lean index aec0305a1a724..3fa0098a8a862 100644 --- a/src/algebra/monoid_algebra/grading.lean +++ b/src/algebra/monoid_algebra/grading.lean @@ -177,7 +177,8 @@ graded_algebra.of_alg_hom _ (decompose_aux f) (begin ext : 2, - dsimp, + simp only [alg_hom.coe_to_monoid_hom, function.comp_app, alg_hom.coe_comp, + function.comp.left_id, alg_hom.coe_id, add_monoid_algebra.of_apply, monoid_hom.coe_comp], rw [decompose_aux_single, direct_sum.coe_alg_hom_of, subtype.coe_mk], end) (λ i x, by rw [decompose_aux_coe f x]) @@ -191,7 +192,7 @@ by apply_instance (direct_sum.decompose (grade_by R f)) := rfl @[simp] lemma grades_by.decompose_single (m : M) (r : R) : - direct_sum.decompose (grade_by R f) (finsupp.single m r) = + direct_sum.decompose (grade_by R f) (finsupp.single m r : add_monoid_algebra R M) = direct_sum.of (λ i : ι, grade_by R f i) (f m) ⟨finsupp.single m r, single_mem_grade_by _ _ _⟩ := decompose_aux_single _ _ _ @@ -205,7 +206,7 @@ by apply_instance @[simp] lemma grade.decompose_single (i : ι) (r : R) : - direct_sum.decompose (grade R : ι → submodule _ _) (finsupp.single i r) = + direct_sum.decompose (grade R : ι → submodule _ _) (finsupp.single i r : add_monoid_algebra _ _) = direct_sum.of (λ i : ι, grade R i) i ⟨finsupp.single i r, single_mem_grade _ _⟩ := decompose_aux_single _ _ _ diff --git a/src/data/set_like/basic.lean b/src/data/set_like/basic.lean index a86e34ae681cf..b83d85eaf1033 100644 --- a/src/data/set_like/basic.lean +++ b/src/data/set_like/basic.lean @@ -82,6 +82,15 @@ Note: if `set_like.coe` is a projection, implementers should create a simp lemma @[simp] lemma mem_carrier {p : my_subobject X} : x ∈ p.carrier ↔ x ∈ (p : set X) := iff.rfl ``` to normalize terms. + +If you declare an unbundled subclass of `set_like`, for example: +``` +class mul_mem_class (S : Type*) (M : Type*) [has_mul M] [set_like S M] where + ... +``` +Then you should *not* repeat the `out_param` declaration, `set_like` will supply the value instead. +This ensures in Lean 4 your subclass will not have issues with synthesis of the `[has_mul M]` +parameter starting before the value of `M` is known. -/ @[protect_proj] class set_like (A : Type*) (B : out_param $ Type*) := diff --git a/src/field_theory/subfield.lean b/src/field_theory/subfield.lean index 1f66d94ff24e0..7d28db29ba330 100644 --- a/src/field_theory/subfield.lean +++ b/src/field_theory/subfield.lean @@ -65,7 +65,7 @@ universes u v w variables {K : Type u} {L : Type v} {M : Type w} [field K] [field L] [field M] /-- `subfield_class S K` states `S` is a type of subsets `s ⊆ K` closed under field operations. -/ -class subfield_class (S : Type*) (K : out_param $ Type*) [field K] [set_like S K] +class subfield_class (S K : Type*) [field K] [set_like S K] extends subring_class S K, inv_mem_class S K : Prop namespace subfield_class diff --git a/src/group_theory/group_action/sub_mul_action.lean b/src/group_theory/group_action/sub_mul_action.lean index 96abf240857d2..34f6cbcf8c2c5 100644 --- a/src/group_theory/group_action/sub_mul_action.lean +++ b/src/group_theory/group_action/sub_mul_action.lean @@ -36,13 +36,23 @@ variables {S : Type u'} {T : Type u''} {R : Type u} {M : Type v} set_option old_structure_cmd true /-- `smul_mem_class S R M` says `S` is a type of subsets `s ≤ M` that are closed under the -scalar action of `R` on `M`. -/ -class smul_mem_class (S : Type*) (R M : out_param $ Type*) [has_smul R M] [set_like S M] := +scalar action of `R` on `M`. + +Note that only `R` is marked as an `out_param` here, since `M` is supplied by the `set_like` +class instead. +-/ +class smul_mem_class (S : Type*) (R : out_param $ Type*) (M : Type*) [has_smul R M] + [set_like S M] := (smul_mem : ∀ {s : S} (r : R) {m : M}, m ∈ s → r • m ∈ s) /-- `vadd_mem_class S R M` says `S` is a type of subsets `s ≤ M` that are closed under the -additive action of `R` on `M`. -/ -class vadd_mem_class (S : Type*) (R M : out_param $ Type*) [has_vadd R M] [set_like S M] := +additive action of `R` on `M`. + +Note that only `R` is marked as an `out_param` here, since `M` is supplied by the `set_like` +class instead. +-/ +class vadd_mem_class (S : Type*) (R : out_param $ Type*) (M : Type*) [has_vadd R M] + [set_like S M] := (vadd_mem : ∀ {s : S} (r : R) {m : M}, m ∈ s → r +ᵥ m ∈ s) attribute [to_additive] smul_mem_class diff --git a/src/group_theory/submonoid/basic.lean b/src/group_theory/submonoid/basic.lean index a503eca7573f6..34d40b22dea98 100644 --- a/src/group_theory/submonoid/basic.lean +++ b/src/group_theory/submonoid/basic.lean @@ -63,13 +63,13 @@ variables [mul_one_class M] {s : set M} variables [add_zero_class A] {t : set A} /-- `one_mem_class S M` says `S` is a type of subsets `s ≤ M`, such that `1 ∈ s` for all `s`. -/ -class one_mem_class (S : Type*) (M : out_param $ Type*) [has_one M] [set_like S M] : Prop := +class one_mem_class (S M : Type*) [has_one M] [set_like S M] : Prop := (one_mem : ∀ (s : S), (1 : M) ∈ s) export one_mem_class (one_mem) /-- `zero_mem_class S M` says `S` is a type of subsets `s ≤ M`, such that `0 ∈ s` for all `s`. -/ -class zero_mem_class (S : Type*) (M : out_param $ Type*) [has_zero M] [set_like S M] : Prop := +class zero_mem_class (S M : Type*) [has_zero M] [set_like S M] : Prop := (zero_mem : ∀ (s : S), (0 : M) ∈ s) export zero_mem_class (zero_mem) @@ -92,7 +92,7 @@ add_decl_doc submonoid.to_subsemigroup /-- `submonoid_class S M` says `S` is a type of subsets `s ≤ M` that contain `1` and are closed under `(*)` -/ -class submonoid_class (S : Type*) (M : out_param $ Type*) [mul_one_class M] [set_like S M] +class submonoid_class (S M : Type*) [mul_one_class M] [set_like S M] extends mul_mem_class S M, one_mem_class S M : Prop section @@ -113,7 +113,7 @@ add_decl_doc add_submonoid.to_add_subsemigroup /-- `add_submonoid_class S M` says `S` is a type of subsets `s ≤ M` that contain `0` and are closed under `(+)` -/ -class add_submonoid_class (S : Type*) (M : out_param $ Type*) [add_zero_class M] [set_like S M] +class add_submonoid_class (S M : Type*) [add_zero_class M] [set_like S M] extends add_mem_class S M, zero_mem_class S M : Prop attribute [to_additive] submonoid submonoid_class diff --git a/src/group_theory/submonoid/pointwise.lean b/src/group_theory/submonoid/pointwise.lean index 4c450f7cbe3f1..c9a53990dacba 100644 --- a/src/group_theory/submonoid/pointwise.lean +++ b/src/group_theory/submonoid/pointwise.lean @@ -441,7 +441,7 @@ begin work_on_goal 1 { intros, apply closure_induction hb, work_on_goal 1 { intros, exact subset_closure ⟨_, _, ‹_›, ‹_›, rfl⟩ } }, all_goals { intros, simp only [mul_zero, zero_mul, zero_mem, - left_distrib, right_distrib, mul_smul_comm, smul_mul_assoc], + left_distrib, right_distrib, mul_smul_comm, smul_mul_assoc]; solve_by_elim [add_mem _ _, zero_mem _] { max_depth := 4, discharger := tactic.interactive.apply_instance } } }, { rw closure_le, rintros _ ⟨a, b, ha, hb, rfl⟩, diff --git a/src/group_theory/subsemigroup/basic.lean b/src/group_theory/subsemigroup/basic.lean index daefae477541c..a380ec63b600b 100644 --- a/src/group_theory/subsemigroup/basic.lean +++ b/src/group_theory/subsemigroup/basic.lean @@ -57,13 +57,13 @@ variables [has_mul M] {s : set M} variables [has_add A] {t : set A} /-- `mul_mem_class S M` says `S` is a type of subsets `s ≤ M` that are closed under `(*)` -/ -class mul_mem_class (S : Type*) (M : out_param $ Type*) [has_mul M] [set_like S M] : Prop := +class mul_mem_class (S M : Type*) [has_mul M] [set_like S M] : Prop := (mul_mem : ∀ {s : S} {a b : M}, a ∈ s → b ∈ s → a * b ∈ s) export mul_mem_class (mul_mem) /-- `add_mem_class S M` says `S` is a type of subsets `s ≤ M` that are closed under `(+)` -/ -class add_mem_class (S : Type*) (M : out_param $ Type*) [has_add M] [set_like S M] : Prop := +class add_mem_class (S M : Type*) [has_add M] [set_like S M] : Prop := (add_mem : ∀ {s : S} {a b : M}, a ∈ s → b ∈ s → a + b ∈ s) export add_mem_class (add_mem) diff --git a/src/logic/basic.lean b/src/logic/basic.lean index 79bafe7763ff2..f834107a1895c 100644 --- a/src/logic/basic.lean +++ b/src/logic/basic.lean @@ -90,7 +90,7 @@ theorem coe_fn_coe_trans /-- Non-dependent version of `coe_fn_coe_trans`, helps `rw` figure out the argument. -/ theorem coe_fn_coe_trans' - {α β γ} {δ : out_param $ _} [has_coe α β] [has_coe_t_aux β γ] [has_coe_to_fun γ (λ _, δ)] + {α β γ} {δ : _} [has_coe α β] [has_coe_t_aux β γ] [has_coe_to_fun γ (λ _, δ)] (x : α) : @coe_fn α _ _ x = @coe_fn β _ _ x := rfl @[simp] theorem coe_fn_coe_base @@ -99,7 +99,7 @@ theorem coe_fn_coe_trans' /-- Non-dependent version of `coe_fn_coe_base`, helps `rw` figure out the argument. -/ theorem coe_fn_coe_base' - {α β} {γ : out_param $ _} [has_coe α β] [has_coe_to_fun β (λ _, γ)] + {α β} {γ : _} [has_coe α β] [has_coe_to_fun β (λ _, γ)] (x : α) : @coe_fn α _ _ x = @coe_fn β _ _ x := rfl -- This instance should have low priority, to ensure we follow the chain diff --git a/src/ring_theory/derivation.lean b/src/ring_theory/derivation.lean index 5b4ad1574a5df..8f050414aca68 100644 --- a/src/ring_theory/derivation.lean +++ b/src/ring_theory/derivation.lean @@ -685,8 +685,8 @@ begin refine ⟨kaehler_differential.one_smul_sub_smul_one_mem_ideal R x, _⟩, apply submodule.subset_span, exact ⟨x, kaehler_differential.D_linear_map_apply R S x⟩ }, - { exact ⟨zero_mem _, zero_mem _⟩ }, - { rintros x y ⟨hx₁, hx₂⟩ ⟨hy₁, hy₂⟩, exact ⟨add_mem hx₁ hy₁, add_mem hx₂ hy₂⟩ }, + { exact ⟨zero_mem _, submodule.zero_mem _⟩ }, + { rintros x y ⟨hx₁, hx₂⟩ ⟨hy₁, hy₂⟩, exact ⟨add_mem hx₁ hy₁, submodule.add_mem _ hx₂ hy₂⟩ }, { rintros r x ⟨hx₁, hx₂⟩, exact ⟨((kaehler_differential.ideal R S).restrict_scalars S).smul_mem r hx₁, submodule.smul_mem _ r hx₂⟩ } end diff --git a/src/ring_theory/non_unital_subsemiring/basic.lean b/src/ring_theory/non_unital_subsemiring/basic.lean index 2d262812f2738..5d85dd07d2a9e 100644 --- a/src/ring_theory/non_unital_subsemiring/basic.lean +++ b/src/ring_theory/non_unital_subsemiring/basic.lean @@ -28,12 +28,12 @@ variables {R : Type u} {S : Type v} {T : Type w} [non_unital_non_assoc_semiring /-- `non_unital_subsemiring_class S R` states that `S` is a type of subsets `s ⊆ R` that are both an additive submonoid and also a multiplicative subsemigroup. -/ -class non_unital_subsemiring_class (S : Type*) (R : out_param $ Type u) +class non_unital_subsemiring_class (S : Type*) (R : Type u) [non_unital_non_assoc_semiring R] [set_like S R] extends add_submonoid_class S R := (mul_mem : ∀ {s : S} {a b : R}, a ∈ s → b ∈ s → a * b ∈ s) @[priority 100] -- See note [lower instance priority] -instance non_unital_subsemiring_class.mul_mem_class (S : Type*) (R : out_param $ Type u) +instance non_unital_subsemiring_class.mul_mem_class (S : Type*) (R : Type u) [non_unital_non_assoc_semiring R] [set_like S R] [h : non_unital_subsemiring_class S R] : mul_mem_class S R := { .. h } diff --git a/src/ring_theory/subring/basic.lean b/src/ring_theory/subring/basic.lean index 54ea6c2f32073..de604180ee48e 100644 --- a/src/ring_theory/subring/basic.lean +++ b/src/ring_theory/subring/basic.lean @@ -71,11 +71,11 @@ section subring_class /-- `subring_class S R` states that `S` is a type of subsets `s ⊆ R` that are both a multiplicative submonoid and an additive subgroup. -/ -class subring_class (S : Type*) (R : out_param $ Type u) [ring R] [set_like S R] +class subring_class (S : Type*) (R : Type u) [ring R] [set_like S R] extends subsemiring_class S R, neg_mem_class S R : Prop @[priority 100] -- See note [lower instance priority] -instance subring_class.add_subgroup_class (S : Type*) (R : out_param $ Type u) [set_like S R] +instance subring_class.add_subgroup_class (S : Type*) (R : Type u) [set_like S R] [ring R] [h : subring_class S R] : add_subgroup_class S R := { .. h } diff --git a/src/ring_theory/subsemiring/basic.lean b/src/ring_theory/subsemiring/basic.lean index f064254ef77bd..7fad17093710a 100644 --- a/src/ring_theory/subsemiring/basic.lean +++ b/src/ring_theory/subsemiring/basic.lean @@ -29,7 +29,7 @@ section add_submonoid_with_one_class /-- `add_submonoid_with_one_class S R` says `S` is a type of subsets `s ≤ R` that contain `0`, `1`, and are closed under `(+)` -/ -class add_submonoid_with_one_class (S : Type*) (R : out_param $ Type*) +class add_submonoid_with_one_class (S : Type*) (R : Type*) [add_monoid_with_one R] [set_like S R] extends add_submonoid_class S R, one_mem_class S R : Prop @@ -55,11 +55,11 @@ section subsemiring_class /-- `subsemiring_class S R` states that `S` is a type of subsets `s ⊆ R` that are both a multiplicative and an additive submonoid. -/ -class subsemiring_class (S : Type*) (R : out_param $ Type u) [non_assoc_semiring R] [set_like S R] +class subsemiring_class (S : Type*) (R : Type u) [non_assoc_semiring R] [set_like S R] extends submonoid_class S R, add_submonoid_class S R : Prop @[priority 100] -- See note [lower instance priority] -instance subsemiring_class.add_submonoid_with_one_class (S : Type*) (R : out_param $ Type u) +instance subsemiring_class.add_submonoid_with_one_class (S : Type*) (R : Type u) [non_assoc_semiring R] [set_like S R] [h : subsemiring_class S R] : add_submonoid_with_one_class S R := { .. h } From f2ce6086713c78a7f880485f7917ea547a215982 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sat, 28 Jan 2023 02:35:05 +0000 Subject: [PATCH 0379/1291] chore(*): rename locale `topological_space` to `topology` (#18293) --- counterexamples/sorgenfrey_line.lean | 2 +- src/analysis/ODE/gronwall.lean | 2 +- src/analysis/ODE/picard_lindelof.lean | 2 +- src/analysis/analytic/basic.lean | 2 +- src/analysis/analytic/composition.lean | 2 +- src/analysis/analytic/inverse.lean | 2 +- src/analysis/analytic/isolated_zeros.lean | 2 +- src/analysis/analytic/linear.lean | 2 +- src/analysis/analytic/radius_liminf.lean | 2 +- src/analysis/analytic/uniqueness.lean | 2 +- .../asymptotics/asymptotic_equivalent.lean | 2 +- src/analysis/asymptotics/asymptotics.lean | 2 +- .../asymptotics/specific_asymptotics.lean | 2 +- .../asymptotics/superpolynomial_decay.lean | 2 +- src/analysis/asymptotics/theta.lean | 2 +- src/analysis/bounded_variation.lean | 2 +- src/analysis/box_integral/basic.lean | 2 +- src/analysis/box_integral/box/basic.lean | 2 +- .../box_integral/box/subbox_induction.lean | 2 +- .../box_integral/divergence_theorem.lean | 2 +- src/analysis/box_integral/integrability.lean | 2 +- .../box_integral/partition/filter.lean | 2 +- .../partition/subbox_induction.lean | 2 +- .../calculus/bump_function_findim.lean | 2 +- src/analysis/calculus/cont_diff.lean | 2 +- src/analysis/calculus/darboux.lean | 2 +- src/analysis/calculus/deriv.lean | 2 +- src/analysis/calculus/diff_cont_on_cl.lean | 2 +- src/analysis/calculus/dslope.lean | 2 +- src/analysis/calculus/extend_deriv.lean | 2 +- src/analysis/calculus/fderiv.lean | 2 +- src/analysis/calculus/fderiv_measurable.lean | 2 +- src/analysis/calculus/fderiv_symmetric.lean | 2 +- .../calculus/formal_multilinear_series.lean | 2 +- src/analysis/calculus/implicit.lean | 2 +- src/analysis/calculus/inverse.lean | 2 +- src/analysis/calculus/iterated_deriv.lean | 2 +- .../calculus/lagrange_multipliers.lean | 2 +- src/analysis/calculus/lhopital.lean | 2 +- src/analysis/calculus/local_extr.lean | 2 +- src/analysis/calculus/mean_value.lean | 2 +- src/analysis/calculus/monotone.lean | 2 +- .../calculus/parametric_integral.lean | 2 +- .../parametric_interval_integral.lean | 2 +- src/analysis/calculus/series.lean | 2 +- src/analysis/calculus/specific_functions.lean | 2 +- src/analysis/calculus/tangent_cone.lean | 2 +- src/analysis/calculus/taylor.lean | 2 +- .../calculus/uniform_limits_deriv.lean | 2 +- src/analysis/complex/abs_max.lean | 2 +- src/analysis/complex/basic.lean | 2 +- src/analysis/complex/cauchy_integral.lean | 2 +- src/analysis/complex/liouville.lean | 2 +- .../complex/locally_uniform_limit.lean | 2 +- src/analysis/complex/open_mapping.lean | 2 +- src/analysis/complex/phragmen_lindelof.lean | 2 +- .../complex/removable_singularity.lean | 2 +- src/analysis/complex/schwarz.lean | 2 +- .../functions_bounded_at_infty.lean | 2 +- .../complex/upper_half_plane/metric.lean | 2 +- .../complex/upper_half_plane/topology.lean | 2 +- src/analysis/convex/extrema.lean | 2 +- src/analysis/convex/integral.lean | 2 +- src/analysis/convex/measure.lean | 2 +- src/analysis/convex/partition_of_unity.lean | 2 +- src/analysis/convolution.lean | 2 +- src/analysis/hofer.lean | 2 +- src/analysis/inner_product_space/basic.lean | 2 +- .../inner_product_space/calculus.lean | 2 +- .../inner_product_space/euclidean_dist.lean | 2 +- .../inner_product_space/l2_space.lean | 2 +- src/analysis/inner_product_space/pi_L2.lean | 2 +- .../inner_product_space/projection.lean | 2 +- src/analysis/locally_convex/abs_convex.lean | 2 +- .../locally_convex/balanced_core_hull.lean | 2 +- src/analysis/locally_convex/basic.lean | 2 +- src/analysis/locally_convex/bounded.lean | 2 +- .../locally_convex/continuous_of_bounded.lean | 2 +- src/analysis/locally_convex/weak_dual.lean | 2 +- .../locally_convex/with_seminorms.lean | 2 +- src/analysis/normed/field/basic.lean | 2 +- src/analysis/normed/group/add_torsor.lean | 2 +- src/analysis/normed/group/basic.lean | 2 +- .../normed/group/controlled_closure.lean | 2 +- src/analysis/normed/group/infinite_sum.lean | 2 +- src/analysis/normed/group/pointwise.lean | 2 +- src/analysis/normed/group/quotient.lean | 2 +- src/analysis/normed/order/basic.lean | 2 +- src/analysis/normed_space/add_torsor.lean | 2 +- src/analysis/normed_space/banach.lean | 2 +- .../normed_space/banach_steinhaus.lean | 2 +- src/analysis/normed_space/basic.lean | 2 +- .../normed_space/bounded_linear_maps.lean | 2 +- .../normed_space/compact_operator.lean | 2 +- src/analysis/normed_space/dual.lean | 2 +- src/analysis/normed_space/exponential.lean | 2 +- .../normed_space/finite_dimension.lean | 2 +- src/analysis/normed_space/lp_space.lean | 2 +- src/analysis/normed_space/multilinear.lean | 2 +- src/analysis/normed_space/operator_norm.lean | 4 ++-- src/analysis/normed_space/pi_Lp.lean | 2 +- src/analysis/normed_space/pointwise.lean | 2 +- src/analysis/normed_space/riesz_lemma.lean | 2 +- src/analysis/normed_space/spectrum.lean | 2 +- src/analysis/normed_space/star/basic.lean | 2 +- src/analysis/normed_space/star/spectrum.lean | 2 +- src/analysis/normed_space/units.lean | 2 +- src/analysis/normed_space/weak_dual.lean | 2 +- src/analysis/p_series.lean | 2 +- src/analysis/seminorm.lean | 2 +- src/analysis/special_functions/arsinh.lean | 2 +- src/analysis/special_functions/bernstein.lean | 2 +- .../special_functions/compare_exp.lean | 2 +- .../special_functions/complex/arg.lean | 2 +- .../special_functions/complex/log.lean | 4 ++-- .../special_functions/complex/log_deriv.lean | 4 ++-- src/analysis/special_functions/exp.lean | 2 +- src/analysis/special_functions/exp_deriv.lean | 2 +- .../special_functions/exponential.lean | 2 +- src/analysis/special_functions/gamma.lean | 2 +- src/analysis/special_functions/gaussian.lean | 2 +- .../special_functions/japanese_bracket.lean | 2 +- src/analysis/special_functions/log/base.lean | 2 +- src/analysis/special_functions/log/basic.lean | 2 +- src/analysis/special_functions/log/deriv.lean | 2 +- .../special_functions/log/monotone.lean | 2 +- .../special_functions/non_integrable.lean | 2 +- .../special_functions/polar_coord.lean | 2 +- .../special_functions/polynomials.lean | 2 +- src/analysis/special_functions/pow.lean | 2 +- src/analysis/special_functions/pow_deriv.lean | 2 +- src/analysis/special_functions/sqrt.lean | 2 +- src/analysis/special_functions/stirling.lean | 2 +- .../trigonometric/arctan.lean | 2 +- .../trigonometric/arctan_deriv.lean | 2 +- .../trigonometric/basic.lean | 2 +- .../trigonometric/complex.lean | 2 +- .../trigonometric/complex_deriv.lean | 2 +- .../trigonometric/deriv.lean | 2 +- .../trigonometric/euler_sine_prod.lean | 2 +- .../trigonometric/inverse.lean | 2 +- .../trigonometric/inverse_deriv.lean | 2 +- src/analysis/specific_limits/basic.lean | 2 +- src/analysis/specific_limits/floor_pow.lean | 2 +- src/analysis/specific_limits/normed.lean | 2 +- src/analysis/subadditive.lean | 2 +- src/combinatorics/additive/behrend.lean | 2 +- .../derangements/exponential.lean | 2 +- src/data/analysis/topology.lean | 2 +- src/data/real/hyperreal.lean | 2 +- src/data/real/pi/leibniz.lean | 2 +- src/data/real/pi/wallis.lean | 2 +- src/data/real/sqrt.lean | 2 +- .../rotation_number/translation_number.lean | 2 +- src/dynamics/ergodic/add_circle.lean | 2 +- src/dynamics/ergodic/conservative.lean | 2 +- src/dynamics/fixed_points/topology.lean | 2 +- src/dynamics/omega_limit.lean | 2 +- src/field_theory/krull_topology.lean | 2 +- src/geometry/manifold/bump_function.lean | 2 +- src/geometry/manifold/charted_space.lean | 2 +- src/geometry/manifold/complex.lean | 2 +- src/geometry/manifold/cont_mdiff.lean | 2 +- src/geometry/manifold/cont_mdiff_mfderiv.lean | 2 +- src/geometry/manifold/diffeomorph.lean | 2 +- .../manifold/local_invariant_properties.lean | 2 +- src/geometry/manifold/mfderiv.lean | 2 +- src/geometry/manifold/partition_of_unity.lean | 2 +- .../smooth_manifold_with_corners.lean | 4 ++-- src/geometry/manifold/tangent_bundle.lean | 2 +- src/geometry/manifold/whitney_embedding.lean | 2 +- .../constructions/borel_space.lean | 2 +- src/measure_theory/constructions/pi.lean | 2 +- src/measure_theory/constructions/polish.lean | 2 +- src/measure_theory/constructions/prod.lean | 2 +- src/measure_theory/covering/besicovitch.lean | 2 +- .../covering/besicovitch_vector_space.lean | 2 +- .../covering/density_theorem.lean | 4 ++-- .../covering/differentiation.lean | 2 +- .../covering/liminf_limsup.lean | 2 +- src/measure_theory/covering/one_dim.lean | 2 +- src/measure_theory/covering/vitali.lean | 2 +- .../covering/vitali_family.lean | 2 +- .../decomposition/unsigned_hahn.lean | 2 +- src/measure_theory/function/ae_eq_fun.lean | 2 +- .../function/ae_eq_of_integral.lean | 2 +- .../conditional_expectation/basic.lean | 2 +- .../conditional_expectation/indicator.lean | 2 +- .../conditional_expectation/real.lean | 2 +- .../function/continuous_map_dense.lean | 2 +- .../function/convergence_in_measure.lean | 2 +- src/measure_theory/function/egorov.lean | 2 +- src/measure_theory/function/jacobian.lean | 2 +- src/measure_theory/function/l1_space.lean | 2 +- .../function/locally_integrable.lean | 2 +- src/measure_theory/function/lp_space.lean | 2 +- .../function/simple_func_dense.lean | 2 +- .../function/simple_func_dense_lp.lean | 2 +- .../function/strongly_measurable/basic.lean | 2 +- .../function/strongly_measurable/lp.lean | 2 +- .../function/uniform_integrable.lean | 2 +- src/measure_theory/group/action.lean | 2 +- src/measure_theory/group/add_circle.lean | 2 +- .../group/fundamental_domain.lean | 2 +- src/measure_theory/group/measure.lean | 2 +- src/measure_theory/integral/average.lean | 2 +- src/measure_theory/integral/bochner.lean | 2 +- .../integral/circle_integral.lean | 2 +- .../integral/divergence_theorem.lean | 2 +- .../integral/integrable_on.lean | 2 +- .../integral/integral_eq_improper.lean | 2 +- .../integral/interval_integral.lean | 2 +- src/measure_theory/integral/lebesgue.lean | 2 +- src/measure_theory/integral/set_integral.lean | 2 +- src/measure_theory/integral/set_to_l1.lean | 2 +- src/measure_theory/measure/doubling.lean | 2 +- .../measure/finite_measure.lean | 2 +- src/measure_theory/measure/haar.lean | 2 +- src/measure_theory/measure/haar_lebesgue.lean | 2 +- src/measure_theory/measure/hausdorff.lean | 2 +- src/measure_theory/measure/lebesgue.lean | 2 +- src/measure_theory/measure/measure_space.lean | 2 +- .../measure/measure_space_def.lean | 2 +- src/measure_theory/measure/open_pos.lean | 2 +- src/measure_theory/measure/outer_measure.lean | 2 +- src/measure_theory/measure/portmanteau.lean | 2 +- .../measure/probability_measure.lean | 2 +- src/measure_theory/measure/regular.lean | 2 +- src/measure_theory/measure/stieltjes.lean | 2 +- .../liouville/liouville_with.lean | 2 +- src/number_theory/liouville/measure.lean | 2 +- src/number_theory/modular_forms/basic.lean | 2 +- src/number_theory/padics/hensel.lean | 2 +- src/number_theory/well_approximable.lean | 2 +- .../filter/zero_and_bounded_at_filter.lean | 2 +- src/probability/borel_cantelli.lean | 2 +- src/probability/ident_distrib.lean | 2 +- .../martingale/borel_cantelli.lean | 2 +- src/probability/martingale/convergence.lean | 2 +- src/probability/martingale/upcrossing.lean | 2 +- src/probability/process/adapted.lean | 2 +- src/probability/process/filtration.lean | 2 +- src/probability/process/hitting_time.lean | 2 +- src/probability/process/stopping.lean | 2 +- src/probability/strong_law.lean | 2 +- src/topology/G_delta.lean | 2 +- src/topology/alexandroff.lean | 2 +- src/topology/algebra/const_mul_action.lean | 2 +- src/topology/algebra/constructions.lean | 2 +- src/topology/algebra/field.lean | 2 +- src/topology/algebra/filter_basis.lean | 2 +- src/topology/algebra/group/basic.lean | 2 +- src/topology/algebra/group/compact.lean | 2 +- src/topology/algebra/group_with_zero.lean | 2 +- src/topology/algebra/infinite_sum.lean | 2 +- src/topology/algebra/module/basic.lean | 2 +- src/topology/algebra/module/linear_pmap.lean | 2 +- .../algebra/module/locally_convex.lean | 2 +- .../algebra/module/strong_topology.lean | 2 +- src/topology/algebra/module/weak_dual.lean | 2 +- src/topology/algebra/monoid.lean | 2 +- src/topology/algebra/mul_action.lean | 2 +- .../algebra/nonarchimedean/adic_topology.lean | 2 +- .../algebra/nonarchimedean/bases.lean | 2 +- src/topology/algebra/open_subgroup.lean | 2 +- src/topology/algebra/order/compact.lean | 2 +- src/topology/algebra/order/extend_from.lean | 2 +- src/topology/algebra/order/extr_closure.lean | 2 +- src/topology/algebra/order/field.lean | 2 +- src/topology/algebra/order/filter.lean | 2 +- src/topology/algebra/order/floor.lean | 2 +- .../algebra/order/intermediate_value.lean | 2 +- src/topology/algebra/order/left_right.lean | 2 +- .../algebra/order/left_right_lim.lean | 2 +- src/topology/algebra/order/liminf_limsup.lean | 4 ++-- .../algebra/order/monotone_continuity.lean | 2 +- .../algebra/order/monotone_convergence.lean | 2 +- src/topology/algebra/order/proj_Icc.lean | 2 +- src/topology/algebra/order/t5.lean | 2 +- src/topology/algebra/ring.lean | 2 +- src/topology/algebra/star.lean | 2 +- src/topology/algebra/uniform_convergence.lean | 2 +- src/topology/algebra/uniform_field.lean | 2 +- src/topology/algebra/uniform_group.lean | 2 +- src/topology/algebra/valuation.lean | 2 +- src/topology/algebra/valued_field.lean | 2 +- src/topology/algebra/with_zero_topology.lean | 2 +- src/topology/bases.lean | 2 +- src/topology/basic.lean | 22 +++++++------------ src/topology/category/Compactum.lean | 2 +- src/topology/compact_open.lean | 2 +- src/topology/connected.lean | 2 +- src/topology/constructions.lean | 2 +- src/topology/continuous_function/bounded.lean | 2 +- src/topology/continuous_function/compact.lean | 2 +- .../stone_weierstrass.lean | 2 +- .../continuous_function/zero_at_infty.lean | 2 +- src/topology/continuous_on.lean | 2 +- src/topology/dense_embedding.lean | 2 +- src/topology/extend_from.lean | 2 +- src/topology/fiber_bundle/basic.lean | 2 +- src/topology/fiber_bundle/constructions.lean | 2 +- src/topology/fiber_bundle/trivialization.lean | 2 +- src/topology/filter.lean | 2 +- src/topology/homeomorph.lean | 2 +- src/topology/homotopy/homotopy_group.lean | 2 +- src/topology/inseparable.lean | 2 +- src/topology/instances/add_circle.lean | 2 +- src/topology/instances/ennreal.lean | 2 +- src/topology/instances/ereal.lean | 2 +- src/topology/instances/irrational.lean | 2 +- src/topology/instances/nnreal.lean | 2 +- src/topology/instances/rat_lemmas.lean | 2 +- src/topology/instances/real.lean | 2 +- src/topology/is_locally_homeomorph.lean | 2 +- src/topology/list.lean | 2 +- src/topology/local_at_target.lean | 2 +- src/topology/local_extr.lean | 2 +- src/topology/local_homeomorph.lean | 2 +- src/topology/locally_constant/basic.lean | 2 +- src/topology/locally_finite.lean | 2 +- src/topology/maps.lean | 2 +- src/topology/metric_space/baire.lean | 2 +- src/topology/metric_space/basic.lean | 2 +- src/topology/metric_space/cau_seq_filter.lean | 2 +- src/topology/metric_space/closeds.lean | 2 +- src/topology/metric_space/completion.lean | 2 +- src/topology/metric_space/contracting.lean | 2 +- .../metric_space/emetric_paracompact.lean | 2 +- src/topology/metric_space/emetric_space.lean | 2 +- src/topology/metric_space/equicontinuity.lean | 2 +- src/topology/metric_space/gluing.lean | 2 +- .../metric_space/gromov_hausdorff.lean | 2 +- .../gromov_hausdorff_realized.lean | 2 +- .../metric_space/hausdorff_dimension.lean | 2 +- .../metric_space/hausdorff_distance.lean | 2 +- src/topology/metric_space/holder.lean | 2 +- src/topology/metric_space/isometry.lean | 2 +- src/topology/metric_space/lipschitz.lean | 2 +- src/topology/metric_space/metrizable.lean | 2 +- .../metric_space/partition_of_unity.lean | 2 +- src/topology/metric_space/pi_nat.lean | 4 ++-- src/topology/metric_space/polish.lean | 2 +- .../metric_space/shrinking_lemma.lean | 2 +- .../metric_space/thickened_indicator.lean | 2 +- src/topology/nhds_set.lean | 4 ++-- src/topology/order.lean | 2 +- src/topology/order/basic.lean | 2 +- src/topology/order/lattice.lean | 2 +- src/topology/paracompact.lean | 2 +- src/topology/partial.lean | 2 +- src/topology/partition_of_unity.lean | 2 +- src/topology/path_connected.lean | 2 +- src/topology/perfect.lean | 2 +- src/topology/semicontinuous.lean | 2 +- src/topology/separation.lean | 2 +- src/topology/sequences.lean | 2 +- src/topology/stone_cech.lean | 2 +- src/topology/subset_properties.lean | 2 +- src/topology/support.lean | 2 +- src/topology/tietze_extension.lean | 2 +- src/topology/uniform_space/basic.lean | 2 +- src/topology/uniform_space/cauchy.lean | 2 +- src/topology/uniform_space/compact.lean | 2 +- .../uniform_space/compact_convergence.lean | 2 +- .../uniform_space/complete_separated.lean | 2 +- src/topology/uniform_space/completion.lean | 2 +- .../uniform_space/equicontinuity.lean | 2 +- src/topology/uniform_space/matrix.lean | 2 +- src/topology/uniform_space/pi.lean | 2 +- src/topology/uniform_space/separation.lean | 2 +- .../uniform_space/uniform_convergence.lean | 2 +- .../uniform_convergence_topology.lean | 2 +- .../uniform_space/uniform_embedding.lean | 2 +- src/topology/unit_interval.lean | 2 +- src/topology/urysohns_lemma.lean | 2 +- 376 files changed, 391 insertions(+), 397 deletions(-) diff --git a/counterexamples/sorgenfrey_line.lean b/counterexamples/sorgenfrey_line.lean index 6b21bb09c031d..f96879bbd74c9 100644 --- a/counterexamples/sorgenfrey_line.lean +++ b/counterexamples/sorgenfrey_line.lean @@ -31,7 +31,7 @@ Prove that the Sorgenfrey line is a paracompact space. -/ open set filter topological_space -open_locale topological_space filter +open_locale topology filter noncomputable theory /-- The Sorgenfrey line. It is the real line with the topology space structure generated by diff --git a/src/analysis/ODE/gronwall.lean b/src/analysis/ODE/gronwall.lean index 7e1fdf7ee5709..04426303cd2c0 100644 --- a/src/analysis/ODE/gronwall.lean +++ b/src/analysis/ODE/gronwall.lean @@ -31,7 +31,7 @@ variables {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] {F : Type*} [normed_add_comm_group F] [normed_space ℝ F] open metric set asymptotics filter real -open_locale classical topological_space nnreal +open_locale classical topology nnreal /-! ### Technical lemmas about `gronwall_bound` -/ diff --git a/src/analysis/ODE/picard_lindelof.lean b/src/analysis/ODE/picard_lindelof.lean index a7bfe958abd54..9d470e9397d30 100644 --- a/src/analysis/ODE/picard_lindelof.lean +++ b/src/analysis/ODE/picard_lindelof.lean @@ -34,7 +34,7 @@ differential equation open filter function set metric topological_space interval_integral measure_theory open measure_theory.measure_space (volume) -open_locale filter topological_space nnreal ennreal nat interval +open_locale filter topology nnreal ennreal nat interval noncomputable theory diff --git a/src/analysis/analytic/basic.lean b/src/analysis/analytic/basic.lean index 3dd9a5016a666..988541d8d5778 100644 --- a/src/analysis/analytic/basic.lean +++ b/src/analysis/analytic/basic.lean @@ -72,7 +72,7 @@ noncomputable theory variables {𝕜 E F G : Type*} -open_locale topological_space classical big_operators nnreal filter ennreal +open_locale topology classical big_operators nnreal filter ennreal open set filter asymptotics namespace formal_multilinear_series diff --git a/src/analysis/analytic/composition.lean b/src/analysis/analytic/composition.lean index eeeea5db90346..87bc85aa02f33 100644 --- a/src/analysis/analytic/composition.lean +++ b/src/analysis/analytic/composition.lean @@ -67,7 +67,7 @@ noncomputable theory variables {𝕜 : Type*} {E F G H : Type*} open filter list -open_locale topological_space big_operators classical nnreal ennreal +open_locale topology big_operators classical nnreal ennreal section topological variables [comm_ring 𝕜] [add_comm_group E] [add_comm_group F] [add_comm_group G] diff --git a/src/analysis/analytic/inverse.lean b/src/analysis/analytic/inverse.lean index daf8eb3717456..f3c4ca2afb134 100644 --- a/src/analysis/analytic/inverse.lean +++ b/src/analysis/analytic/inverse.lean @@ -26,7 +26,7 @@ we prove that they coincide and study their properties (notably convergence). -/ -open_locale big_operators classical topological_space +open_locale big_operators classical topology open finset filter namespace formal_multilinear_series diff --git a/src/analysis/analytic/isolated_zeros.lean b/src/analysis/analytic/isolated_zeros.lean index aef178e90decb..c311ba2b0ad6a 100644 --- a/src/analysis/analytic/isolated_zeros.lean +++ b/src/analysis/analytic/isolated_zeros.lean @@ -30,7 +30,7 @@ useful in this setup. open_locale classical open filter function nat formal_multilinear_series emetric set -open_locale topological_space big_operators +open_locale topology big_operators variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] {E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] {s : E} diff --git a/src/analysis/analytic/linear.lean b/src/analysis/analytic/linear.lean index 7d88f847c567f..e0c124303f844 100644 --- a/src/analysis/analytic/linear.lean +++ b/src/analysis/analytic/linear.lean @@ -17,7 +17,7 @@ variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] {F : Type*} [normed_add_comm_group F] [normed_space 𝕜 F] {G : Type*} [normed_add_comm_group G] [normed_space 𝕜 G] -open_locale topological_space classical big_operators nnreal ennreal +open_locale topology classical big_operators nnreal ennreal open set filter asymptotics noncomputable theory diff --git a/src/analysis/analytic/radius_liminf.lean b/src/analysis/analytic/radius_liminf.lean index f706addfdb95b..bd745bebaee50 100644 --- a/src/analysis/analytic/radius_liminf.lean +++ b/src/analysis/analytic/radius_liminf.lean @@ -17,7 +17,7 @@ variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] {E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] {F : Type*} [normed_add_comm_group F] [normed_space 𝕜 F] -open_locale topological_space classical big_operators nnreal ennreal +open_locale topology classical big_operators nnreal ennreal open filter asymptotics namespace formal_multilinear_series diff --git a/src/analysis/analytic/uniqueness.lean b/src/analysis/analytic/uniqueness.lean index 65596baccb9b3..57a5268a74db5 100644 --- a/src/analysis/analytic/uniqueness.lean +++ b/src/analysis/analytic/uniqueness.lean @@ -19,7 +19,7 @@ variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] {F : Type*} [normed_add_comm_group F] [normed_space 𝕜 F] open set -open_locale topological_space ennreal +open_locale topology ennreal namespace analytic_on diff --git a/src/analysis/asymptotics/asymptotic_equivalent.lean b/src/analysis/asymptotics/asymptotic_equivalent.lean index e0dc4078fc0e5..5d50ec2ce1c0b 100644 --- a/src/analysis/asymptotics/asymptotic_equivalent.lean +++ b/src/analysis/asymptotics/asymptotic_equivalent.lean @@ -57,7 +57,7 @@ This is to enable `calc` support, as `calc` requires that the last two explicit namespace asymptotics open filter function -open_locale topological_space +open_locale topology section normed_add_comm_group diff --git a/src/analysis/asymptotics/asymptotics.lean b/src/analysis/asymptotics/asymptotics.lean index edf8db8b19103..21eff942ac3c3 100644 --- a/src/analysis/asymptotics/asymptotics.lean +++ b/src/analysis/asymptotics/asymptotics.lean @@ -45,7 +45,7 @@ the Fréchet derivative.) -/ open filter set -open_locale topological_space big_operators classical filter nnreal +open_locale topology big_operators classical filter nnreal namespace asymptotics diff --git a/src/analysis/asymptotics/specific_asymptotics.lean b/src/analysis/asymptotics/specific_asymptotics.lean index ef8af93edbb0a..33909a5bc401e 100644 --- a/src/analysis/asymptotics/specific_asymptotics.lean +++ b/src/analysis/asymptotics/specific_asymptotics.lean @@ -14,7 +14,7 @@ theory developped in `analysis.asymptotics.asymptotics`. -/ open filter asymptotics -open_locale topological_space +open_locale topology section normed_field diff --git a/src/analysis/asymptotics/superpolynomial_decay.lean b/src/analysis/asymptotics/superpolynomial_decay.lean index 4a9f43f38c4e6..bf461b0a82afe 100644 --- a/src/analysis/asymptotics/superpolynomial_decay.lean +++ b/src/analysis/asymptotics/superpolynomial_decay.lean @@ -46,7 +46,7 @@ https://ncatlab.org/nlab/show/rapidly+decreasing+function namespace asymptotics -open_locale topological_space polynomial +open_locale topology polynomial open filter /-- `f` has superpolynomial decay in parameter `k` along filter `l` if diff --git a/src/analysis/asymptotics/theta.lean b/src/analysis/asymptotics/theta.lean index 7d4f9070084a4..fb18642e27344 100644 --- a/src/analysis/asymptotics/theta.lean +++ b/src/analysis/asymptotics/theta.lean @@ -13,7 +13,7 @@ In this file we define `asymptotics.is_Theta l f g` (notation: `f =Θ[l] g`) as -/ open filter -open_locale topological_space +open_locale topology namespace asymptotics diff --git a/src/analysis/bounded_variation.lean b/src/analysis/bounded_variation.lean index 5809e1b74a97b..7e74573ea7701 100644 --- a/src/analysis/bounded_variation.lean +++ b/src/analysis/bounded_variation.lean @@ -45,7 +45,7 @@ it possible to use the complete linear order structure of `ℝ≥0∞`. The proo more tedious with an `ℝ`-valued or `ℝ≥0`-valued variation, since one would always need to check that the sets one uses are nonempty and bounded above as these are only conditionally complete. -/ -open_locale big_operators nnreal ennreal topological_space uniform_convergence +open_locale big_operators nnreal ennreal topology uniform_convergence open set measure_theory filter variables {α β : Type*} [linear_order α] [linear_order β] diff --git a/src/analysis/box_integral/basic.lean b/src/analysis/box_integral/basic.lean index 971bd2284073c..eb617857d3890 100644 --- a/src/analysis/box_integral/basic.lean +++ b/src/analysis/box_integral/basic.lean @@ -48,7 +48,7 @@ non-Riemann filter (e.g., Henstock-Kurzweil and McShane). integral -/ -open_locale big_operators classical topological_space nnreal filter uniformity box_integral +open_locale big_operators classical topology nnreal filter uniformity box_integral open set finset function filter metric box_integral.integration_params noncomputable theory diff --git a/src/analysis/box_integral/box/basic.lean b/src/analysis/box_integral/box/basic.lean index 74828f30cec50..e9c525883e371 100644 --- a/src/analysis/box_integral/box/basic.lean +++ b/src/analysis/box_integral/box/basic.lean @@ -52,7 +52,7 @@ rectangular box open set function metric filter noncomputable theory -open_locale nnreal classical topological_space +open_locale nnreal classical topology namespace box_integral diff --git a/src/analysis/box_integral/box/subbox_induction.lean b/src/analysis/box_integral/box/subbox_induction.lean index a5f63b8179bdc..1f25231c4836d 100644 --- a/src/analysis/box_integral/box/subbox_induction.lean +++ b/src/analysis/box_integral/box/subbox_induction.lean @@ -27,7 +27,7 @@ rectangular box, induction -/ open set finset function filter metric -open_locale classical topological_space filter ennreal +open_locale classical topology filter ennreal noncomputable theory namespace box_integral diff --git a/src/analysis/box_integral/divergence_theorem.lean b/src/analysis/box_integral/divergence_theorem.lean index 85b72e111e843..97600b5182095 100644 --- a/src/analysis/box_integral/divergence_theorem.lean +++ b/src/analysis/box_integral/divergence_theorem.lean @@ -38,7 +38,7 @@ Henstock-Kurzweil integral. Henstock-Kurzweil integral, integral, Stokes theorem, divergence theorem -/ -open_locale classical big_operators nnreal ennreal topological_space box_integral +open_locale classical big_operators nnreal ennreal topology box_integral open continuous_linear_map (lsmul) filter set finset metric box_integral.integration_params (GP GP_le) noncomputable theory diff --git a/src/analysis/box_integral/integrability.lean b/src/analysis/box_integral/integrability.lean index 0cd3e3289fc46..eba1bdd4c83d7 100644 --- a/src/analysis/box_integral/integrability.lean +++ b/src/analysis/box_integral/integrability.lean @@ -18,7 +18,7 @@ Henstock and `⊥` integrable) with the same integral. The proof is based on integral, McShane integral, Bochner integral -/ -open_locale classical nnreal ennreal topological_space big_operators +open_locale classical nnreal ennreal topology big_operators universes u v diff --git a/src/analysis/box_integral/partition/filter.lean b/src/analysis/box_integral/partition/filter.lean index 72e36eae9ec13..7f33e1c02fede 100644 --- a/src/analysis/box_integral/partition/filter.lean +++ b/src/analysis/box_integral/partition/filter.lean @@ -162,7 +162,7 @@ integral, rectangular box, partition, filter -/ open set function filter metric finset bool -open_locale classical topological_space filter nnreal +open_locale classical topology filter nnreal noncomputable theory namespace box_integral diff --git a/src/analysis/box_integral/partition/subbox_induction.lean b/src/analysis/box_integral/partition/subbox_induction.lean index e171fc752b88e..230dbddcfe4a3 100644 --- a/src/analysis/box_integral/partition/subbox_induction.lean +++ b/src/analysis/box_integral/partition/subbox_induction.lean @@ -29,7 +29,7 @@ partition, tagged partition, Henstock integral namespace box_integral open set metric -open_locale classical topological_space +open_locale classical topology noncomputable theory variables {ι : Type*} [fintype ι] {I J : box ι} diff --git a/src/analysis/calculus/bump_function_findim.lean b/src/analysis/calculus/bump_function_findim.lean index dc35767b374ca..bd39456517103 100644 --- a/src/analysis/calculus/bump_function_findim.lean +++ b/src/analysis/calculus/bump_function_findim.lean @@ -24,7 +24,7 @@ noncomputable theory open set metric topological_space function asymptotics measure_theory finite_dimensional continuous_linear_map filter measure_theory.measure -open_locale pointwise topological_space nnreal big_operators convolution +open_locale pointwise topology nnreal big_operators convolution variables {E : Type*} [normed_add_comm_group E] diff --git a/src/analysis/calculus/cont_diff.lean b/src/analysis/calculus/cont_diff.lean index 68f3a7110ce41..3f8f95d40cdf6 100644 --- a/src/analysis/calculus/cont_diff.lean +++ b/src/analysis/calculus/cont_diff.lean @@ -170,7 +170,7 @@ local attribute [instance, priority 1001] normed_add_comm_group.to_add_comm_group normed_space.to_module' add_comm_group.to_add_comm_monoid open set fin filter function -open_locale topological_space +open_locale topology variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] {E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] diff --git a/src/analysis/calculus/darboux.lean b/src/analysis/calculus/darboux.lean index 32445699fdd8b..1d25068eee2f2 100644 --- a/src/analysis/calculus/darboux.lean +++ b/src/analysis/calculus/darboux.lean @@ -14,7 +14,7 @@ intermediate values. The proof is based on the -/ open filter set -open_locale topological_space classical +open_locale topology classical variables {a b : ℝ} {f f' : ℝ → ℝ} diff --git a/src/analysis/calculus/deriv.lean b/src/analysis/calculus/deriv.lean index 0db145990c9b4..1665fcc1ecb7c 100644 --- a/src/analysis/calculus/deriv.lean +++ b/src/analysis/calculus/deriv.lean @@ -83,7 +83,7 @@ See the explanations there. universes u v w noncomputable theory -open_locale classical topological_space big_operators filter ennreal polynomial +open_locale classical topology big_operators filter ennreal polynomial open filter asymptotics set open continuous_linear_map (smul_right smul_right_one_eq_iff) diff --git a/src/analysis/calculus/diff_cont_on_cl.lean b/src/analysis/calculus/diff_cont_on_cl.lean index 9dac40f6a3871..2ca22622a955b 100644 --- a/src/analysis/calculus/diff_cont_on_cl.lean +++ b/src/analysis/calculus/diff_cont_on_cl.lean @@ -14,7 +14,7 @@ this property and prove basic facts about this predicate. -/ open set filter metric -open_locale topological_space +open_locale topology variables (𝕜 : Type*) {E F G : Type*} [nontrivially_normed_field 𝕜] [normed_add_comm_group E] [normed_add_comm_group F] [normed_space 𝕜 E] [normed_space 𝕜 F] [normed_add_comm_group G] diff --git a/src/analysis/calculus/dslope.lean b/src/analysis/calculus/dslope.lean index 2f20481ada53d..8899e73779af2 100644 --- a/src/analysis/calculus/dslope.lean +++ b/src/analysis/calculus/dslope.lean @@ -17,7 +17,7 @@ In this file we define `dslope` and prove some basic lemmas about its continuity differentiability. -/ -open_locale classical topological_space filter +open_locale classical topology filter open function set filter variables {𝕜 E : Type*} [nontrivially_normed_field 𝕜] [normed_add_comm_group E] [normed_space 𝕜 E] diff --git a/src/analysis/calculus/extend_deriv.lean b/src/analysis/calculus/extend_deriv.lean index 4e07fb7c59b69..a80b6895a99a2 100644 --- a/src/analysis/calculus/extend_deriv.lean +++ b/src/analysis/calculus/extend_deriv.lean @@ -24,7 +24,7 @@ variables {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] {F : Type*} [normed_add_comm_group F] [normed_space ℝ F] open filter set metric continuous_linear_map -open_locale topological_space +open_locale topology local attribute [mono] prod_mono /-- If a function `f` is differentiable in a convex open set and continuous on its closure, and its diff --git a/src/analysis/calculus/fderiv.lean b/src/analysis/calculus/fderiv.lean index 2dcaaae1573f0..c2baacbaf4115 100644 --- a/src/analysis/calculus/fderiv.lean +++ b/src/analysis/calculus/fderiv.lean @@ -114,7 +114,7 @@ derivative, differentiable, Fréchet, calculus -/ open filter asymptotics continuous_linear_map set metric -open_locale topological_space classical nnreal filter asymptotics ennreal +open_locale topology classical nnreal filter asymptotics ennreal noncomputable theory diff --git a/src/analysis/calculus/fderiv_measurable.lean b/src/analysis/calculus/fderiv_measurable.lean index 4a4a653356c2b..6c52ff3f17c76 100644 --- a/src/analysis/calculus/fderiv_measurable.lean +++ b/src/analysis/calculus/fderiv_measurable.lean @@ -76,7 +76,7 @@ noncomputable theory open set metric asymptotics filter continuous_linear_map open topological_space (second_countable_topology) measure_theory -open_locale topological_space +open_locale topology namespace continuous_linear_map diff --git a/src/analysis/calculus/fderiv_symmetric.lean b/src/analysis/calculus/fderiv_symmetric.lean index 34f00e42edab5..70a4aeab6a8ff 100644 --- a/src/analysis/calculus/fderiv_symmetric.lean +++ b/src/analysis/calculus/fderiv_symmetric.lean @@ -48,7 +48,7 @@ rectangle are contained in `s` by convexity. The general case follows by lineari -/ open asymptotics set -open_locale topological_space +open_locale topology variables {E F : Type*} [normed_add_comm_group E] [normed_space ℝ E] [normed_add_comm_group F] [normed_space ℝ F] diff --git a/src/analysis/calculus/formal_multilinear_series.lean b/src/analysis/calculus/formal_multilinear_series.lean index 4c7400f75d276..36c44df149a6e 100644 --- a/src/analysis/calculus/formal_multilinear_series.lean +++ b/src/analysis/calculus/formal_multilinear_series.lean @@ -25,7 +25,7 @@ multilinear, formal series noncomputable theory open set fin -open_locale topological_space +open_locale topology variables {𝕜 𝕜' E F G : Type*} diff --git a/src/analysis/calculus/implicit.lean b/src/analysis/calculus/implicit.lean index 2a50efb0c3174..9e4b3ecc5dca0 100644 --- a/src/analysis/calculus/implicit.lean +++ b/src/analysis/calculus/implicit.lean @@ -45,7 +45,7 @@ implicit function, inverse function noncomputable theory -open_locale topological_space +open_locale topology open filter open continuous_linear_map (fst snd smul_right ker_prod) open continuous_linear_equiv (of_bijective) diff --git a/src/analysis/calculus/inverse.lean b/src/analysis/calculus/inverse.lean index 3ef859903176b..fa25e0d52e932 100644 --- a/src/analysis/calculus/inverse.lean +++ b/src/analysis/calculus/inverse.lean @@ -57,7 +57,7 @@ derivative, strictly differentiable, continuously differentiable, smooth, invers -/ open function set filter metric -open_locale topological_space classical nnreal +open_locale topology classical nnreal noncomputable theory diff --git a/src/analysis/calculus/iterated_deriv.lean b/src/analysis/calculus/iterated_deriv.lean index 57637f6b3ffd9..0282cb9a4197a 100644 --- a/src/analysis/calculus/iterated_deriv.lean +++ b/src/analysis/calculus/iterated_deriv.lean @@ -41,7 +41,7 @@ iterated Fréchet derivative. -/ noncomputable theory -open_locale classical topological_space big_operators +open_locale classical topology big_operators open filter asymptotics set diff --git a/src/analysis/calculus/lagrange_multipliers.lean b/src/analysis/calculus/lagrange_multipliers.lean index 6d3f314434da8..98273ea1853e9 100644 --- a/src/analysis/calculus/lagrange_multipliers.lean +++ b/src/analysis/calculus/lagrange_multipliers.lean @@ -27,7 +27,7 @@ lagrange multiplier, local extremum -/ open filter set -open_locale topological_space filter big_operators +open_locale topology filter big_operators variables {E F : Type*} [normed_add_comm_group E] [normed_space ℝ E] [complete_space E] [normed_add_comm_group F] [normed_space ℝ F] [complete_space F] {f : E → F} {φ : E → ℝ} {x₀ : E} {f' : E →L[ℝ] F} {φ' : E →L[ℝ] ℝ} diff --git a/src/analysis/calculus/lhopital.lean b/src/analysis/calculus/lhopital.lean index 9ed6d0d6911ac..200fd62607420 100644 --- a/src/analysis/calculus/lhopital.lean +++ b/src/analysis/calculus/lhopital.lean @@ -29,7 +29,7 @@ L'Hôpital's rule, L'Hopital's rule -/ open filter set -open_locale filter topological_space pointwise +open_locale filter topology pointwise variables {a b : ℝ} (hab : a < b) {l : filter ℝ} {f f' g g' : ℝ → ℝ} diff --git a/src/analysis/calculus/local_extr.lean b/src/analysis/calculus/local_extr.lean index 6694ec4d17df0..70ee18d3fc286 100644 --- a/src/analysis/calculus/local_extr.lean +++ b/src/analysis/calculus/local_extr.lean @@ -65,7 +65,7 @@ local extremum, Fermat's Theorem, Rolle's Theorem universes u v open filter set -open_locale topological_space classical polynomial +open_locale topology classical polynomial section module diff --git a/src/analysis/calculus/mean_value.lean b/src/analysis/calculus/mean_value.lean index 53a1238a6040b..4238fa2836b40 100644 --- a/src/analysis/calculus/mean_value.lean +++ b/src/analysis/calculus/mean_value.lean @@ -69,7 +69,7 @@ variables {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] {F : Type*} [normed_add_comm_group F] [normed_space ℝ F] open metric set asymptotics continuous_linear_map filter -open_locale classical topological_space nnreal +open_locale classical topology nnreal /-! ### One-dimensional fencing inequalities -/ diff --git a/src/analysis/calculus/monotone.lean b/src/analysis/calculus/monotone.lean index 1bc2b999d16ad..cc6205136f304 100644 --- a/src/analysis/calculus/monotone.lean +++ b/src/analysis/calculus/monotone.lean @@ -31,7 +31,7 @@ behavior of `μ [x, y]`. -/ open set filter function metric measure_theory measure_theory.measure is_doubling_measure -open_locale topological_space +open_locale topology /-- If `(f y - f x) / (y - x)` converges to a limit as `y` tends to `x`, then the same goes if `y` is shifted a little bit, i.e., `f (y + (y-x)^2) - f x) / (y - x)` converges to the same limit. diff --git a/src/analysis/calculus/parametric_integral.lean b/src/analysis/calculus/parametric_integral.lean index 0ba29b2f0fba4..ea58e3346dfc1 100644 --- a/src/analysis/calculus/parametric_integral.lean +++ b/src/analysis/calculus/parametric_integral.lean @@ -54,7 +54,7 @@ integral, derivative noncomputable theory open topological_space measure_theory filter metric -open_locale topological_space filter +open_locale topology filter variables {α : Type*} [measurable_space α] {μ : measure α} {𝕜 : Type*} [is_R_or_C 𝕜] {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] [normed_space 𝕜 E] diff --git a/src/analysis/calculus/parametric_interval_integral.lean b/src/analysis/calculus/parametric_interval_integral.lean index f6e82e4cef673..807cd3f5df161 100644 --- a/src/analysis/calculus/parametric_interval_integral.lean +++ b/src/analysis/calculus/parametric_interval_integral.lean @@ -14,7 +14,7 @@ integrals. -/ open topological_space measure_theory filter metric -open_locale topological_space filter interval +open_locale topology filter interval variables {𝕜 : Type*} [is_R_or_C 𝕜] {μ : measure ℝ} {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] [normed_space 𝕜 E] diff --git a/src/analysis/calculus/series.lean b/src/analysis/calculus/series.lean index bc0e56143e4bc..a5028c8f9f57b 100644 --- a/src/analysis/calculus/series.lean +++ b/src/analysis/calculus/series.lean @@ -22,7 +22,7 @@ We also give versions of these statements which are localized to a set. -/ open set metric topological_space function asymptotics filter -open_locale topological_space nnreal big_operators +open_locale topology nnreal big_operators variables {α β 𝕜 E F : Type*} [is_R_or_C 𝕜] diff --git a/src/analysis/calculus/specific_functions.lean b/src/analysis/calculus/specific_functions.lean index b75744a009566..7cca302310603 100644 --- a/src/analysis/calculus/specific_functions.lean +++ b/src/analysis/calculus/specific_functions.lean @@ -46,7 +46,7 @@ function cannot have: -/ noncomputable theory -open_locale classical topological_space +open_locale classical topology open polynomial real filter set function open_locale polynomial diff --git a/src/analysis/calculus/tangent_cone.lean b/src/analysis/calculus/tangent_cone.lean index 9e8ab6413dd15..10009a260f271 100644 --- a/src/analysis/calculus/tangent_cone.lean +++ b/src/analysis/calculus/tangent_cone.lean @@ -33,7 +33,7 @@ properties of the tangent cone we prove here. variables (𝕜 : Type*) [nontrivially_normed_field 𝕜] open filter set -open_locale topological_space +open_locale topology section tangent_cone diff --git a/src/analysis/calculus/taylor.lean b/src/analysis/calculus/taylor.lean index 31527b5521359..2ac9483bbf23c 100644 --- a/src/analysis/calculus/taylor.lean +++ b/src/analysis/calculus/taylor.lean @@ -42,7 +42,7 @@ Taylor polynomial, Taylor's theorem -/ -open_locale big_operators interval topological_space nat +open_locale big_operators interval topology nat open set variables {𝕜 E F : Type*} diff --git a/src/analysis/calculus/uniform_limits_deriv.lean b/src/analysis/calculus/uniform_limits_deriv.lean index 0e1a612cfc533..ae111d0325927 100644 --- a/src/analysis/calculus/uniform_limits_deriv.lean +++ b/src/analysis/calculus/uniform_limits_deriv.lean @@ -93,7 +93,7 @@ uniform convergence, limits of derivatives -/ open filter -open_locale uniformity filter topological_space +open_locale uniformity filter topology section limits_of_derivatives diff --git a/src/analysis/complex/abs_max.lean b/src/analysis/complex/abs_max.lean index 61a9eae8571e3..d99498d5098f4 100644 --- a/src/analysis/complex/abs_max.lean +++ b/src/analysis/complex/abs_max.lean @@ -77,7 +77,7 @@ maximum modulus principle, complex analysis -/ open topological_space metric set filter asymptotics function measure_theory affine_map -open_locale topological_space filter nnreal real +open_locale topology filter nnreal real universes u v w variables {E : Type u} [normed_add_comm_group E] [normed_space ℂ E] diff --git a/src/analysis/complex/basic.lean b/src/analysis/complex/basic.lean index e33c4d037cd5f..13819d109a3b3 100644 --- a/src/analysis/complex/basic.lean +++ b/src/analysis/complex/basic.lean @@ -37,7 +37,7 @@ noncomputable theory namespace complex -open_locale complex_conjugate topological_space +open_locale complex_conjugate topology instance : has_norm ℂ := ⟨abs⟩ diff --git a/src/analysis/complex/cauchy_integral.lean b/src/analysis/complex/cauchy_integral.lean index 6a123025b7ec6..d285e42c300b9 100644 --- a/src/analysis/complex/cauchy_integral.lean +++ b/src/analysis/complex/cauchy_integral.lean @@ -143,7 +143,7 @@ Cauchy-Goursat theorem, Cauchy integral formula -/ open topological_space set measure_theory interval_integral metric filter function -open_locale interval real nnreal ennreal topological_space big_operators +open_locale interval real nnreal ennreal topology big_operators noncomputable theory diff --git a/src/analysis/complex/liouville.lean b/src/analysis/complex/liouville.lean index 95b25ddbf2e5a..e297de7e96bd2 100644 --- a/src/analysis/complex/liouville.lean +++ b/src/analysis/complex/liouville.lean @@ -21,7 +21,7 @@ The proof is based on the Cauchy integral formula for the derivative of an analy -/ open topological_space metric set filter asymptotics function measure_theory -open_locale topological_space filter nnreal real +open_locale topology filter nnreal real universes u v variables {E : Type u} [normed_add_comm_group E] [normed_space ℂ E] diff --git a/src/analysis/complex/locally_uniform_limit.lean b/src/analysis/complex/locally_uniform_limit.lean index 549d05693862d..36e01d439ce67 100644 --- a/src/analysis/complex/locally_uniform_limit.lean +++ b/src/analysis/complex/locally_uniform_limit.lean @@ -21,7 +21,7 @@ subset of the complex plane. -/ open set metric measure_theory filter complex interval_integral -open_locale real topological_space +open_locale real topology variables {E ι : Type*} [normed_add_comm_group E] [normed_space ℂ E] [complete_space E] {U K : set ℂ} {z : ℂ} {M r δ : ℝ} {φ : filter ι} {F : ι → ℂ → E} {f g : ℂ → E} diff --git a/src/analysis/complex/open_mapping.lean b/src/analysis/complex/open_mapping.lean index 17819086e6986..bb944e46be3a0 100644 --- a/src/analysis/complex/open_mapping.lean +++ b/src/analysis/complex/open_mapping.lean @@ -30,7 +30,7 @@ second step is implemented in `diff_cont_on_cl.ball_subset_image_closed_ball`. -/ open set filter metric complex -open_locale topological_space +open_locale topology variables {E : Type*} [normed_add_comm_group E] [normed_space ℂ E] {U : set E} {f : ℂ → ℂ} {g : E → ℂ} {z₀ w : ℂ} {ε r m : ℝ} diff --git a/src/analysis/complex/phragmen_lindelof.lean b/src/analysis/complex/phragmen_lindelof.lean index 4ed88cff0bf7b..71e462a2b384f 100644 --- a/src/analysis/complex/phragmen_lindelof.lean +++ b/src/analysis/complex/phragmen_lindelof.lean @@ -43,7 +43,7 @@ real plane has only finitely many limit cycles). -/ open set function filter asymptotics metric complex -open_locale topological_space filter real +open_locale topology filter real local notation `expR` := real.exp diff --git a/src/analysis/complex/removable_singularity.lean b/src/analysis/complex/removable_singularity.lean index 8ce0db71f46c4..2e52005a61a26 100644 --- a/src/analysis/complex/removable_singularity.lean +++ b/src/analysis/complex/removable_singularity.lean @@ -17,7 +17,7 @@ function `function.update f c (lim (𝓝[≠] c) f)` is complex differentiable i -/ open topological_space metric set filter asymptotics function -open_locale topological_space filter nnreal real +open_locale topology filter nnreal real universe u variables {E : Type u} [normed_add_comm_group E] [normed_space ℂ E] [complete_space E] diff --git a/src/analysis/complex/schwarz.lean b/src/analysis/complex/schwarz.lean index 39c31bbd89302..96f06468ded00 100644 --- a/src/analysis/complex/schwarz.lean +++ b/src/analysis/complex/schwarz.lean @@ -47,7 +47,7 @@ Schwarz lemma -/ open metric set function filter topological_space -open_locale topological_space +open_locale topology namespace complex diff --git a/src/analysis/complex/upper_half_plane/functions_bounded_at_infty.lean b/src/analysis/complex/upper_half_plane/functions_bounded_at_infty.lean index 3a6ab5170dd87..53538b4fa044d 100644 --- a/src/analysis/complex/upper_half_plane/functions_bounded_at_infty.lean +++ b/src/analysis/complex/upper_half_plane/functions_bounded_at_infty.lean @@ -19,7 +19,7 @@ Both of which are relevant for defining modular forms. open complex filter -open_locale topological_space upper_half_plane +open_locale topology upper_half_plane noncomputable theory diff --git a/src/analysis/complex/upper_half_plane/metric.lean b/src/analysis/complex/upper_half_plane/metric.lean index efb5f7f75c981..2811716915ee1 100644 --- a/src/analysis/complex/upper_half_plane/metric.lean +++ b/src/analysis/complex/upper_half_plane/metric.lean @@ -24,7 +24,7 @@ ball/sphere with another center and radius. noncomputable theory -open_locale upper_half_plane complex_conjugate nnreal topological_space +open_locale upper_half_plane complex_conjugate nnreal topology open set metric filter real variables {z w : ℍ} {r R : ℝ} diff --git a/src/analysis/complex/upper_half_plane/topology.lean b/src/analysis/complex/upper_half_plane/topology.lean index 65fba4c6df834..418466831de71 100644 --- a/src/analysis/complex/upper_half_plane/topology.lean +++ b/src/analysis/complex/upper_half_plane/topology.lean @@ -19,7 +19,7 @@ various instances. noncomputable theory open set filter function topological_space complex -open_locale filter topological_space upper_half_plane +open_locale filter topology upper_half_plane namespace upper_half_plane diff --git a/src/analysis/convex/extrema.lean b/src/analysis/convex/extrema.lean index 340be610d6774..c5ee65948f40c 100644 --- a/src/analysis/convex/extrema.lean +++ b/src/analysis/convex/extrema.lean @@ -21,7 +21,7 @@ variables {E β : Type*} [add_comm_group E] [topological_space E] {s : set E} open set filter function -open_locale classical topological_space +open_locale classical topology /-- Helper lemma for the more general case: `is_min_on.of_is_local_min_on_of_convex_on`. diff --git a/src/analysis/convex/integral.lean b/src/analysis/convex/integral.lean index 2bb94caed20b2..2ae4b1cb63efc 100644 --- a/src/analysis/convex/integral.lean +++ b/src/analysis/convex/integral.lean @@ -35,7 +35,7 @@ convex, integral, center mass, average value, Jensen's inequality -/ open measure_theory measure_theory.measure metric set filter topological_space function -open_locale topological_space big_operators ennreal convex +open_locale topology big_operators ennreal convex variables {α E F : Type*} {m0 : measurable_space α} [normed_add_comm_group E] [normed_space ℝ E] [complete_space E] diff --git a/src/analysis/convex/measure.lean b/src/analysis/convex/measure.lean index 3a82f562c9094..e2366037f7134 100644 --- a/src/analysis/convex/measure.lean +++ b/src/analysis/convex/measure.lean @@ -16,7 +16,7 @@ convex set in `E`. Then the frontier of `s` has measure zero (see `convex.add_ha -/ open measure_theory measure_theory.measure set metric filter finite_dimensional (finrank) -open_locale topological_space nnreal ennreal +open_locale topology nnreal ennreal variables {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] [measurable_space E] [borel_space E] [finite_dimensional ℝ E] (μ : measure E) [is_add_haar_measure μ] {s : set E} diff --git a/src/analysis/convex/partition_of_unity.lean b/src/analysis/convex/partition_of_unity.lean index 8096dcc4ae1bf..de3e06952c621 100644 --- a/src/analysis/convex/partition_of_unity.lean +++ b/src/analysis/convex/partition_of_unity.lean @@ -25,7 +25,7 @@ partition of unity -/ open set function -open_locale big_operators topological_space +open_locale big_operators topology variables {ι X E : Type*} [topological_space X] [add_comm_group E] [module ℝ E] diff --git a/src/analysis/convolution.lean b/src/analysis/convolution.lean index 44f2aab93c6ff..4cb07c26156f9 100644 --- a/src/analysis/convolution.lean +++ b/src/analysis/convolution.lean @@ -88,7 +88,7 @@ The following notations are localized in the locale `convolution`: open set function filter measure_theory measure_theory.measure topological_space open continuous_linear_map metric -open_locale pointwise topological_space nnreal filter +open_locale pointwise topology nnreal filter universes u𝕜 uG uE uE' uE'' uF uF' uF'' uP diff --git a/src/analysis/hofer.lean b/src/analysis/hofer.lean index 8212d269731c0..eee0d642da7ed 100644 --- a/src/analysis/hofer.lean +++ b/src/analysis/hofer.lean @@ -18,7 +18,7 @@ example of a proof needing to construct a sequence by induction in the middle of * H. Hofer and C. Viterbo, *The Weinstein conjecture in the presence of holomorphic spheres* -/ -open_locale classical topological_space big_operators +open_locale classical topology big_operators open filter finset local notation `d` := dist diff --git a/src/analysis/inner_product_space/basic.lean b/src/analysis/inner_product_space/basic.lean index 2ac8ac71ba5b1..9684da3de7d98 100644 --- a/src/analysis/inner_product_space/basic.lean +++ b/src/analysis/inner_product_space/basic.lean @@ -70,7 +70,7 @@ The Coq code is available at the following address: - π` and `(log x).im ≤ π`. `log 0 = 0`-/ @@ -188,7 +188,7 @@ end complex section log_deriv open complex filter -open_locale topological_space +open_locale topology variables {α : Type*} diff --git a/src/analysis/special_functions/complex/log_deriv.lean b/src/analysis/special_functions/complex/log_deriv.lean index 99ed7f3c415ae..2d43df3418d50 100644 --- a/src/analysis/special_functions/complex/log_deriv.lean +++ b/src/analysis/special_functions/complex/log_deriv.lean @@ -17,7 +17,7 @@ namespace complex open set filter -open_locale real topological_space +open_locale real topology /-- `complex.exp` as a `local_homeomorph` with `source = {z | -π < im z < π}` and `target = {z | 0 < re z} ∪ {z | im z ≠ 0}`. This definition is used to prove that `complex.log` @@ -66,7 +66,7 @@ end complex section log_deriv open complex filter -open_locale topological_space +open_locale topology variables {α : Type*} [topological_space α] {E : Type*} [normed_add_comm_group E] [normed_space ℂ E] diff --git a/src/analysis/special_functions/exp.lean b/src/analysis/special_functions/exp.lean index 553a2a3e86118..0a00ddc6869ed 100644 --- a/src/analysis/special_functions/exp.lean +++ b/src/analysis/special_functions/exp.lean @@ -22,7 +22,7 @@ exp noncomputable theory open finset filter metric asymptotics set function -open_locale classical topological_space +open_locale classical topology namespace complex diff --git a/src/analysis/special_functions/exp_deriv.lean b/src/analysis/special_functions/exp_deriv.lean index 07d2b5e7aae93..8bd19a0140f00 100644 --- a/src/analysis/special_functions/exp_deriv.lean +++ b/src/analysis/special_functions/exp_deriv.lean @@ -20,7 +20,7 @@ exp, derivative noncomputable theory open filter asymptotics set function -open_locale classical topological_space +open_locale classical topology namespace complex variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] [normed_algebra 𝕜 ℂ] diff --git a/src/analysis/special_functions/exponential.lean b/src/analysis/special_functions/exponential.lean index 1fdf68e1062d6..cc3a4dcc3fc7d 100644 --- a/src/analysis/special_functions/exponential.lean +++ b/src/analysis/special_functions/exponential.lean @@ -45,7 +45,7 @@ We prove most result for an arbitrary field `𝕂`, and then specialize to `𝕂 -/ open filter is_R_or_C continuous_multilinear_map normed_field asymptotics -open_locale nat topological_space big_operators ennreal +open_locale nat topology big_operators ennreal section any_field_any_algebra diff --git a/src/analysis/special_functions/gamma.lean b/src/analysis/special_functions/gamma.lean index 5a56c0228d5c1..58018959827f1 100644 --- a/src/analysis/special_functions/gamma.lean +++ b/src/analysis/special_functions/gamma.lean @@ -52,7 +52,7 @@ Gamma noncomputable theory open filter interval_integral set real measure_theory asymptotics -open_locale nat topological_space ennreal big_operators complex_conjugate +open_locale nat topology ennreal big_operators complex_conjugate lemma integral_exp_neg_Ioi : ∫ (x : ℝ) in Ioi 0, exp (-x) = 1 := begin diff --git a/src/analysis/special_functions/gaussian.lean b/src/analysis/special_functions/gaussian.lean index 9eabecd37a2f4..ca021f84a0bb9 100644 --- a/src/analysis/special_functions/gaussian.lean +++ b/src/analysis/special_functions/gaussian.lean @@ -21,7 +21,7 @@ We prove various versions of the formula for the Gaussian integral: noncomputable theory open real set measure_theory filter asymptotics -open_locale real topological_space +open_locale real topology open complex (hiding exp continuous_exp abs_of_nonneg) notation `cexp` := complex.exp diff --git a/src/analysis/special_functions/japanese_bracket.lean b/src/analysis/special_functions/japanese_bracket.lean index 30ab3a86d24a3..1ebd195adc4bc 100644 --- a/src/analysis/special_functions/japanese_bracket.lean +++ b/src/analysis/special_functions/japanese_bracket.lean @@ -27,7 +27,7 @@ than the dimension. noncomputable theory -open_locale big_operators nnreal filter topological_space ennreal +open_locale big_operators nnreal filter topology ennreal open asymptotics filter set real measure_theory finite_dimensional diff --git a/src/analysis/special_functions/log/base.lean b/src/analysis/special_functions/log/base.lean index 032f344c49c4d..a4d25840c2688 100644 --- a/src/analysis/special_functions/log/base.lean +++ b/src/analysis/special_functions/log/base.lean @@ -23,7 +23,7 @@ logarithm, continuity -/ open set filter function -open_locale topological_space +open_locale topology noncomputable theory namespace real diff --git a/src/analysis/special_functions/log/basic.lean b/src/analysis/special_functions/log/basic.lean index 6667ebe175a9b..a02cc7d13b992 100644 --- a/src/analysis/special_functions/log/basic.lean +++ b/src/analysis/special_functions/log/basic.lean @@ -21,7 +21,7 @@ logarithm, continuity -/ open set filter function -open_locale topological_space +open_locale topology noncomputable theory namespace real diff --git a/src/analysis/special_functions/log/deriv.lean b/src/analysis/special_functions/log/deriv.lean index 434ae117622d4..23629e5f7f7c5 100644 --- a/src/analysis/special_functions/log/deriv.lean +++ b/src/analysis/special_functions/log/deriv.lean @@ -19,7 +19,7 @@ logarithm, derivative -/ open filter finset set -open_locale topological_space big_operators +open_locale topology big_operators namespace real diff --git a/src/analysis/special_functions/log/monotone.lean b/src/analysis/special_functions/log/monotone.lean index 665632df01b50..97e9a4f0d190f 100644 --- a/src/analysis/special_functions/log/monotone.lean +++ b/src/analysis/special_functions/log/monotone.lean @@ -18,7 +18,7 @@ logarithm, tonality -/ open set filter function -open_locale topological_space +open_locale topology noncomputable theory namespace real diff --git a/src/analysis/special_functions/non_integrable.lean b/src/analysis/special_functions/non_integrable.lean index fcb9d043b7207..a762d8bb77764 100644 --- a/src/analysis/special_functions/non_integrable.lean +++ b/src/analysis/special_functions/non_integrable.lean @@ -36,7 +36,7 @@ latter lemma to prove that the function `λ x, x⁻¹` is integrable on `a..b` i integrable function -/ -open_locale measure_theory topological_space interval nnreal ennreal +open_locale measure_theory topology interval nnreal ennreal open measure_theory topological_space set filter asymptotics interval_integral variables {E F : Type*} [normed_add_comm_group E] [normed_space ℝ E] [second_countable_topology E] diff --git a/src/analysis/special_functions/polar_coord.lean b/src/analysis/special_functions/polar_coord.lean index 81c941f14e955..05cc1d283346e 100644 --- a/src/analysis/special_functions/polar_coord.lean +++ b/src/analysis/special_functions/polar_coord.lean @@ -19,7 +19,7 @@ It satisfies the following change of variables formula (see `integral_comp_polar noncomputable theory open real set measure_theory -open_locale real topological_space +open_locale real topology /-- The polar coordinates local homeomorphism in `ℝ^2`, mapping `(r cos θ, r sin θ)` to `(r, θ)`. It is a homeomorphism between `ℝ^2 - (-∞, 0]` and `(0, +∞) × (-π, π)`. -/ diff --git a/src/analysis/special_functions/polynomials.lean b/src/analysis/special_functions/polynomials.lean index b9c190472cf71..41b3c39a452b1 100644 --- a/src/analysis/special_functions/polynomials.lean +++ b/src/analysis/special_functions/polynomials.lean @@ -21,7 +21,7 @@ polynomials. -/ open filter finset asymptotics -open_locale asymptotics polynomial topological_space +open_locale asymptotics polynomial topology namespace polynomial diff --git a/src/analysis/special_functions/pow.lean b/src/analysis/special_functions/pow.lean index aa46f782dbe80..5570344bf000c 100644 --- a/src/analysis/special_functions/pow.lean +++ b/src/analysis/special_functions/pow.lean @@ -20,7 +20,7 @@ We also prove basic properties of these functions. noncomputable theory -open_locale classical real topological_space nnreal ennreal filter big_operators complex_conjugate +open_locale classical real topology nnreal ennreal filter big_operators complex_conjugate open filter finset set namespace complex diff --git a/src/analysis/special_functions/pow_deriv.lean b/src/analysis/special_functions/pow_deriv.lean index 8b39292431f8d..6793eae1134f3 100644 --- a/src/analysis/special_functions/pow_deriv.lean +++ b/src/analysis/special_functions/pow_deriv.lean @@ -18,7 +18,7 @@ We also prove differentiability and provide derivatives for the power functions noncomputable theory -open_locale classical real topological_space nnreal ennreal filter +open_locale classical real topology nnreal ennreal filter open filter namespace complex diff --git a/src/analysis/special_functions/sqrt.lean b/src/analysis/special_functions/sqrt.lean index 6304b60ea46f0..3288842d449ea 100644 --- a/src/analysis/special_functions/sqrt.lean +++ b/src/analysis/special_functions/sqrt.lean @@ -17,7 +17,7 @@ sqrt, differentiable -/ open set -open_locale topological_space +open_locale topology namespace real diff --git a/src/analysis/special_functions/stirling.lean b/src/analysis/special_functions/stirling.lean index 1e4fcd66f44b4..8e859206a3df4 100644 --- a/src/analysis/special_functions/stirling.lean +++ b/src/analysis/special_functions/stirling.lean @@ -31,7 +31,7 @@ and prove that $a = \sqrt{\pi}$. Here the main ingredient is the convergence of formula for `π`. -/ -open_locale topological_space real big_operators nat +open_locale topology real big_operators nat open finset filter nat real namespace stirling diff --git a/src/analysis/special_functions/trigonometric/arctan.lean b/src/analysis/special_functions/trigonometric/arctan.lean index ab95d5755caab..10ce7f424c9fc 100644 --- a/src/analysis/special_functions/trigonometric/arctan.lean +++ b/src/analysis/special_functions/trigonometric/arctan.lean @@ -17,7 +17,7 @@ noncomputable theory namespace real open set filter -open_locale topological_space real +open_locale topology real lemma tan_add {x y : ℝ} (h : ((∀ k : ℤ, x ≠ (2 * k + 1) * π / 2) ∧ ∀ l : ℤ, y ≠ (2 * l + 1) * π / 2) diff --git a/src/analysis/special_functions/trigonometric/arctan_deriv.lean b/src/analysis/special_functions/trigonometric/arctan_deriv.lean index 7f46be4b3f99d..d36c11cf4d688 100644 --- a/src/analysis/special_functions/trigonometric/arctan_deriv.lean +++ b/src/analysis/special_functions/trigonometric/arctan_deriv.lean @@ -17,7 +17,7 @@ noncomputable theory namespace real open set filter -open_locale topological_space real +open_locale topology real lemma has_strict_deriv_at_tan {x : ℝ} (h : cos x ≠ 0) : has_strict_deriv_at tan (1 / (cos x)^2) x := diff --git a/src/analysis/special_functions/trigonometric/basic.lean b/src/analysis/special_functions/trigonometric/basic.lean index 1c923a79a3099..f605746cd1007 100644 --- a/src/analysis/special_functions/trigonometric/basic.lean +++ b/src/analysis/special_functions/trigonometric/basic.lean @@ -40,7 +40,7 @@ sin, cos, tan, angle -/ noncomputable theory -open_locale classical topological_space filter +open_locale classical topology filter open set filter namespace complex diff --git a/src/analysis/special_functions/trigonometric/complex.lean b/src/analysis/special_functions/trigonometric/complex.lean index bbcf7ee514b77..2e10bbd8af22a 100644 --- a/src/analysis/special_functions/trigonometric/complex.lean +++ b/src/analysis/special_functions/trigonometric/complex.lean @@ -138,7 +138,7 @@ lemma tan_eq {z : ℂ} tan z = (tan z.re + tanh z.im * I) / (1 - tan z.re * tanh z.im * I) := by convert tan_add_mul_I h; exact (re_add_im z).symm -open_locale topological_space +open_locale topology lemma continuous_on_tan : continuous_on tan {x | cos x ≠ 0} := continuous_on_sin.div continuous_on_cos $ λ x, id diff --git a/src/analysis/special_functions/trigonometric/complex_deriv.lean b/src/analysis/special_functions/trigonometric/complex_deriv.lean index a8a78b04f7acb..6e2fcb65e2b18 100644 --- a/src/analysis/special_functions/trigonometric/complex_deriv.lean +++ b/src/analysis/special_functions/trigonometric/complex_deriv.lean @@ -31,7 +31,7 @@ lemma has_deriv_at_tan {x : ℂ} (h : cos x ≠ 0) : has_deriv_at tan (1 / (cos x)^2) x := (has_strict_deriv_at_tan h).has_deriv_at -open_locale topological_space +open_locale topology lemma tendsto_abs_tan_of_cos_eq_zero {x : ℂ} (hx : cos x = 0) : tendsto (λ x, abs (tan x)) (𝓝[≠] x) at_top := diff --git a/src/analysis/special_functions/trigonometric/deriv.lean b/src/analysis/special_functions/trigonometric/deriv.lean index 30bf58d909902..e6bf6f493a075 100644 --- a/src/analysis/special_functions/trigonometric/deriv.lean +++ b/src/analysis/special_functions/trigonometric/deriv.lean @@ -22,7 +22,7 @@ sin, cos, tan, angle -/ noncomputable theory -open_locale classical topological_space filter +open_locale classical topology filter open set filter namespace complex diff --git a/src/analysis/special_functions/trigonometric/euler_sine_prod.lean b/src/analysis/special_functions/trigonometric/euler_sine_prod.lean index 8563d8909fec5..21fb9a9b5eb98 100644 --- a/src/analysis/special_functions/trigonometric/euler_sine_prod.lean +++ b/src/analysis/special_functions/trigonometric/euler_sine_prod.lean @@ -20,7 +20,7 @@ is to prove a recurrence relation for the integrals `∫ x in 0..π/2, cos 2 z x generalising the arguments used to prove Wallis' limit formula for `π`. -/ -open_locale real topological_space big_operators +open_locale real topology big_operators open real set filter interval_integral measure_theory.measure_space namespace euler_sine diff --git a/src/analysis/special_functions/trigonometric/inverse.lean b/src/analysis/special_functions/trigonometric/inverse.lean index 2112d3833abbe..81ce4c01ad38c 100644 --- a/src/analysis/special_functions/trigonometric/inverse.lean +++ b/src/analysis/special_functions/trigonometric/inverse.lean @@ -17,7 +17,7 @@ Basic inequalities on trigonometric functions. -/ noncomputable theory -open_locale classical topological_space filter +open_locale classical topology filter open set filter open_locale real diff --git a/src/analysis/special_functions/trigonometric/inverse_deriv.lean b/src/analysis/special_functions/trigonometric/inverse_deriv.lean index 152623840582c..bdff998aeb4d3 100644 --- a/src/analysis/special_functions/trigonometric/inverse_deriv.lean +++ b/src/analysis/special_functions/trigonometric/inverse_deriv.lean @@ -13,7 +13,7 @@ Derivatives of `arcsin` and `arccos`. -/ noncomputable theory -open_locale classical topological_space filter +open_locale classical topology filter open set filter open_locale real diff --git a/src/analysis/specific_limits/basic.lean b/src/analysis/specific_limits/basic.lean index 6b908a43b83c8..2ec78fcb19695 100644 --- a/src/analysis/specific_limits/basic.lean +++ b/src/analysis/specific_limits/basic.lean @@ -19,7 +19,7 @@ instances of these such as `ℝ`, `ℝ≥0` and `ℝ≥0∞`. noncomputable theory open classical set function filter finset metric -open_locale classical topological_space nat big_operators uniformity nnreal ennreal +open_locale classical topology nat big_operators uniformity nnreal ennreal variables {α : Type*} {β : Type*} {ι : Type*} diff --git a/src/analysis/specific_limits/floor_pow.lean b/src/analysis/specific_limits/floor_pow.lean index 22a9af21cf370..9ef025f425f28 100644 --- a/src/analysis/specific_limits/floor_pow.lean +++ b/src/analysis/specific_limits/floor_pow.lean @@ -19,7 +19,7 @@ We state several auxiliary results pertaining to sequences of the form `⌊c^n -/ open filter finset -open_locale topological_space big_operators +open_locale topology big_operators /-- If a monotone sequence `u` is such that `u n / n` tends to a limit `l` along subsequences with exponential growth rate arbitrarily close to `1`, then `u n / n` tends to `l`. -/ diff --git a/src/analysis/specific_limits/normed.lean b/src/analysis/specific_limits/normed.lean index 47c7d7a5d578a..830407e577d14 100644 --- a/src/analysis/specific_limits/normed.lean +++ b/src/analysis/specific_limits/normed.lean @@ -19,7 +19,7 @@ spaces. noncomputable theory open classical set function filter finset metric asymptotics -open_locale classical topological_space nat big_operators uniformity nnreal ennreal +open_locale classical topology nat big_operators uniformity nnreal ennreal variables {α : Type*} {β : Type*} {ι : Type*} diff --git a/src/analysis/subadditive.lean b/src/analysis/subadditive.lean index afda32b190afa..6739d92a06530 100644 --- a/src/analysis/subadditive.lean +++ b/src/analysis/subadditive.lean @@ -17,7 +17,7 @@ convenience). This result is known as Fekete's lemma in the literature. noncomputable theory open set filter -open_locale topological_space +open_locale topology /-- A real-valued sequence is subadditive if it satisfies the inequality `u (m + n) ≤ u m + u n` for all `m, n`. -/ diff --git a/src/combinatorics/additive/behrend.lean b/src/combinatorics/additive/behrend.lean index aa84a203c8664..a2c1ac3cd379e 100644 --- a/src/combinatorics/additive/behrend.lean +++ b/src/combinatorics/additive/behrend.lean @@ -251,7 +251,7 @@ begin rwa cast_ne_zero, end -open_locale filter topological_space +open_locale filter topology open real section numerical_bounds diff --git a/src/combinatorics/derangements/exponential.lean b/src/combinatorics/derangements/exponential.lean index afd91585d11dd..6ab726aa8a1ce 100644 --- a/src/combinatorics/derangements/exponential.lean +++ b/src/combinatorics/derangements/exponential.lean @@ -16,7 +16,7 @@ The specific lemma is `num_derangements_tendsto_inv_e`. open filter open_locale big_operators -open_locale topological_space +open_locale topology theorem num_derangements_tendsto_inv_e : tendsto (λ n, (num_derangements n : ℝ) / n.factorial) at_top diff --git a/src/data/analysis/topology.lean b/src/data/analysis/topology.lean index 36a161a35f71c..e3e37b4705d22 100644 --- a/src/data/analysis/topology.lean +++ b/src/data/analysis/topology.lean @@ -21,7 +21,7 @@ This file provides infrastructure to compute with topological spaces. open set open filter (hiding realizer) -open_locale topological_space +open_locale topology /-- A `ctop α σ` is a realization of a topology (basis) on `α`, represented by a type `σ` together with operations for the top element and diff --git a/src/data/real/hyperreal.lean b/src/data/real/hyperreal.lean index 5f8f7ae1e1ace..054a124f0077a 100644 --- a/src/data/real/hyperreal.lean +++ b/src/data/real/hyperreal.lean @@ -11,7 +11,7 @@ import analysis.specific_limits.basic -/ open filter filter.germ -open_locale topological_space classical +open_locale topology classical /-- Hyperreal numbers on the ultrafilter extending the cofinite filter -/ @[derive [linear_ordered_field, inhabited]] diff --git a/src/data/real/pi/leibniz.lean b/src/data/real/pi/leibniz.lean index 5ba42465a6d9d..be7ce817b362d 100644 --- a/src/data/real/pi/leibniz.lean +++ b/src/data/real/pi/leibniz.lean @@ -10,7 +10,7 @@ import analysis.special_functions.trigonometric.arctan_deriv namespace real open filter set -open_locale classical big_operators topological_space real +open_locale classical big_operators topology real local notation (name := abs) `|`x`|` := abs x /-- This theorem establishes **Leibniz's series for `π`**: The alternating sum of the reciprocals diff --git a/src/data/real/pi/wallis.lean b/src/data/real/pi/wallis.lean index 7eb62616e6fb5..c58e97b4aad7c 100644 --- a/src/data/real/pi/wallis.lean +++ b/src/data/real/pi/wallis.lean @@ -32,7 +32,7 @@ algebraic manipulation. * `real.tendsto_prod_pi_div_two`: the Wallis product formula. -/ -open_locale real topological_space big_operators nat +open_locale real topology big_operators nat open filter finset interval_integral namespace real diff --git a/src/data/real/sqrt.lean b/src/data/real/sqrt.lean index d2cfa45030bf5..27e44303b38c8 100644 --- a/src/data/real/sqrt.lean +++ b/src/data/real/sqrt.lean @@ -34,7 +34,7 @@ square root -/ open set filter -open_locale filter nnreal topological_space +open_locale filter nnreal topology namespace nnreal diff --git a/src/dynamics/circle/rotation_number/translation_number.lean b/src/dynamics/circle/rotation_number/translation_number.lean index 221512b2b7ca8..0851e653cab39 100644 --- a/src/dynamics/circle/rotation_number/translation_number.lean +++ b/src/dynamics/circle/rotation_number/translation_number.lean @@ -116,7 +116,7 @@ circle homeomorphism, rotation number -/ open filter set function (hiding commute) int -open_locale topological_space classical +open_locale topology classical /-! ### Definition and monoid structure diff --git a/src/dynamics/ergodic/add_circle.lean b/src/dynamics/ergodic/add_circle.lean index a1c40c7ea404e..7db687edc7e6c 100644 --- a/src/dynamics/ergodic/add_circle.lean +++ b/src/dynamics/ergodic/add_circle.lean @@ -27,7 +27,7 @@ This file contains proofs of ergodicity for maps of the additive circle. -/ open set function measure_theory measure_theory.measure filter metric -open_locale measure_theory nnreal ennreal topological_space pointwise +open_locale measure_theory nnreal ennreal topology pointwise namespace add_circle diff --git a/src/dynamics/ergodic/conservative.lean b/src/dynamics/ergodic/conservative.lean index 0835c3ff44a00..94280422b3db5 100644 --- a/src/dynamics/ergodic/conservative.lean +++ b/src/dynamics/ergodic/conservative.lean @@ -38,7 +38,7 @@ conservative dynamical system, Poincare recurrence theorem noncomputable theory open classical set filter measure_theory finset function topological_space -open_locale classical topological_space +open_locale classical topology variables {ι : Type*} {α : Type*} [measurable_space α] {f : α → α} {s : set α} {μ : measure α} diff --git a/src/dynamics/fixed_points/topology.lean b/src/dynamics/fixed_points/topology.lean index 665a645538719..88ce01ec5c485 100644 --- a/src/dynamics/fixed_points/topology.lean +++ b/src/dynamics/fixed_points/topology.lean @@ -22,7 +22,7 @@ fixed points, iterates variables {α : Type*} [topological_space α] [t2_space α] {f : α → α} open function filter -open_locale topological_space +open_locale topology /-- If the iterates `f^[n] x` converge to `y` and `f` is continuous at `y`, then `y` is a fixed point for `f`. -/ diff --git a/src/dynamics/omega_limit.lean b/src/dynamics/omega_limit.lean index 0e95a6320dcb6..c6be612795c50 100644 --- a/src/dynamics/omega_limit.lean +++ b/src/dynamics/omega_limit.lean @@ -30,7 +30,7 @@ endowed with an order. -/ open set function filter -open_locale topological_space +open_locale topology /-! ### Definition and notation diff --git a/src/field_theory/krull_topology.lean b/src/field_theory/krull_topology.lean index 79eff4c09f6de..03c6438783c48 100644 --- a/src/field_theory/krull_topology.lean +++ b/src/field_theory/krull_topology.lean @@ -199,7 +199,7 @@ group_filter_basis.is_topological_group (gal_group_basis K L) section krull_t2 -open_locale topological_space filter +open_locale topology filter /-- Let `L/E/K` be a tower of fields with `E/K` finite. Then `Gal(L/E)` is an open subgroup of `L ≃ₐ[K] L`. -/ diff --git a/src/geometry/manifold/bump_function.lean b/src/geometry/manifold/bump_function.lean index 425e72abc444a..999bcb4927999 100644 --- a/src/geometry/manifold/bump_function.lean +++ b/src/geometry/manifold/bump_function.lean @@ -33,7 +33,7 @@ variables {M : Type uM} [topological_space M] [charted_space H M] [smooth_manifold_with_corners I M] open function filter finite_dimensional set -open_locale topological_space manifold classical filter big_operators +open_locale topology manifold classical filter big_operators noncomputable theory diff --git a/src/geometry/manifold/charted_space.lean b/src/geometry/manifold/charted_space.lean index 3edfdbbcac523..9584af912c1cb 100644 --- a/src/geometry/manifold/charted_space.lean +++ b/src/geometry/manifold/charted_space.lean @@ -108,7 +108,7 @@ composition of local equivs with `≫`. -/ noncomputable theory -open_locale classical topological_space +open_locale classical topology open filter universes u diff --git a/src/geometry/manifold/complex.lean b/src/geometry/manifold/complex.lean index bbd2e7ddfff65..5e3d546fa2a71 100644 --- a/src/geometry/manifold/complex.lean +++ b/src/geometry/manifold/complex.lean @@ -36,7 +36,7 @@ stalks, such as the Weierstrass preparation theorem. -/ -open_locale manifold topological_space +open_locale manifold topology open complex namespace mdifferentiable diff --git a/src/geometry/manifold/cont_mdiff.lean b/src/geometry/manifold/cont_mdiff.lean index e96446cc2f90e..66382c546e2e2 100644 --- a/src/geometry/manifold/cont_mdiff.lean +++ b/src/geometry/manifold/cont_mdiff.lean @@ -43,7 +43,7 @@ in terms of extended charts in `cont_mdiff_on_iff` and `cont_mdiff_iff`. -/ open set function filter charted_space smooth_manifold_with_corners -open_locale topological_space manifold +open_locale topology manifold /-! ### Definition of smooth functions between manifolds -/ diff --git a/src/geometry/manifold/cont_mdiff_mfderiv.lean b/src/geometry/manifold/cont_mdiff_mfderiv.lean index 97174c760d71b..7ef6055e3c23f 100644 --- a/src/geometry/manifold/cont_mdiff_mfderiv.lean +++ b/src/geometry/manifold/cont_mdiff_mfderiv.lean @@ -21,7 +21,7 @@ and related notions. -/ open set function filter charted_space smooth_manifold_with_corners -open_locale topological_space manifold +open_locale topology manifold /-! ### Definition of smooth functions between manifolds -/ diff --git a/src/geometry/manifold/diffeomorph.lean b/src/geometry/manifold/diffeomorph.lean index 244e0eaa36750..a02d3e5dc4d45 100644 --- a/src/geometry/manifold/diffeomorph.lean +++ b/src/geometry/manifold/diffeomorph.lean @@ -43,7 +43,7 @@ practice. diffeomorphism, manifold -/ -open_locale manifold topological_space +open_locale manifold topology open function set variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] diff --git a/src/geometry/manifold/local_invariant_properties.lean b/src/geometry/manifold/local_invariant_properties.lean index e1e770aac570a..2ba32fbfda0dd 100644 --- a/src/geometry/manifold/local_invariant_properties.lean +++ b/src/geometry/manifold/local_invariant_properties.lean @@ -43,7 +43,7 @@ in the one for `lift_prop_within_at`. -/ noncomputable theory -open_locale classical manifold topological_space +open_locale classical manifold topology open set filter diff --git a/src/geometry/manifold/mfderiv.lean b/src/geometry/manifold/mfderiv.lean index ee0b19a28d9ad..557a2322c06b3 100644 --- a/src/geometry/manifold/mfderiv.lean +++ b/src/geometry/manifold/mfderiv.lean @@ -93,7 +93,7 @@ Derivative, manifold -/ noncomputable theory -open_locale classical topological_space manifold +open_locale classical topology manifold open set diff --git a/src/geometry/manifold/partition_of_unity.lean b/src/geometry/manifold/partition_of_unity.lean index 64cd831e553e0..7598e00e1f479 100644 --- a/src/geometry/manifold/partition_of_unity.lean +++ b/src/geometry/manifold/partition_of_unity.lean @@ -58,7 +58,7 @@ smooth bump function, partition of unity universes uι uE uH uM uF open function filter finite_dimensional set -open_locale topological_space manifold classical filter big_operators +open_locale topology manifold classical filter big_operators noncomputable theory diff --git a/src/geometry/manifold/smooth_manifold_with_corners.lean b/src/geometry/manifold/smooth_manifold_with_corners.lean index 45cd85e25fc39..067ccf65e4a29 100644 --- a/src/geometry/manifold/smooth_manifold_with_corners.lean +++ b/src/geometry/manifold/smooth_manifold_with_corners.lean @@ -114,7 +114,7 @@ noncomputable theory universes u v w u' v' w' open set filter function -open_locale manifold filter topological_space +open_locale manifold filter topology localized "notation (name := with_top.nat.top) `∞` := (⊤ : ℕ∞)" in manifold @@ -702,7 +702,7 @@ instance : smooth_manifold_with_corners I s := { ..s.has_groupoid (cont_diff_gro end topological_space.opens section extended_charts -open_locale topological_space +open_locale topology variables {𝕜 E M H E' M' H' : Type*} [nontrivially_normed_field 𝕜] [normed_add_comm_group E] [normed_space 𝕜 E] [topological_space H] [topological_space M] diff --git a/src/geometry/manifold/tangent_bundle.lean b/src/geometry/manifold/tangent_bundle.lean index 1624cce93995e..a761b141522da 100644 --- a/src/geometry/manifold/tangent_bundle.lean +++ b/src/geometry/manifold/tangent_bundle.lean @@ -76,7 +76,7 @@ noncomputable theory universe u open topological_space set -open_locale manifold topological_space +open_locale manifold topology /-- Core structure used to create a smooth bundle above `M` (a manifold over the model with corner `I`) with fiber the normed vector space `F` over `𝕜`, which is trivial in the chart domains diff --git a/src/geometry/manifold/whitney_embedding.lean b/src/geometry/manifold/whitney_embedding.lean index 5bd93d1a83745..02240f02a1f1a 100644 --- a/src/geometry/manifold/whitney_embedding.lean +++ b/src/geometry/manifold/whitney_embedding.lean @@ -32,7 +32,7 @@ variables {ι : Type uι} {M : Type uM} [topological_space M] [charted_space H M] [smooth_manifold_with_corners I M] open function filter finite_dimensional set -open_locale topological_space manifold classical filter big_operators +open_locale topology manifold classical filter big_operators noncomputable theory diff --git a/src/measure_theory/constructions/borel_space.lean b/src/measure_theory/constructions/borel_space.lean index 4d7319d3a66ad..295b8f72dacfc 100644 --- a/src/measure_theory/constructions/borel_space.lean +++ b/src/measure_theory/constructions/borel_space.lean @@ -46,7 +46,7 @@ import topology.metric_space.metrizable noncomputable theory open classical set filter measure_theory -open_locale classical big_operators topological_space nnreal ennreal measure_theory +open_locale classical big_operators topology nnreal ennreal measure_theory universes u v w x y variables {α β γ γ₂ δ : Type*} {ι : Sort y} {s t u : set α} diff --git a/src/measure_theory/constructions/pi.lean b/src/measure_theory/constructions/pi.lean index 4ced0f4c76403..b2c120dbd7656 100644 --- a/src/measure_theory/constructions/pi.lean +++ b/src/measure_theory/constructions/pi.lean @@ -53,7 +53,7 @@ finitary product measure noncomputable theory open function set measure_theory.outer_measure filter measurable_space encodable -open_locale classical big_operators topological_space ennreal +open_locale classical big_operators topology ennreal universes u v diff --git a/src/measure_theory/constructions/polish.lean b/src/measure_theory/constructions/polish.lean index def85aa002d32..0e0655101819d 100644 --- a/src/measure_theory/constructions/polish.lean +++ b/src/measure_theory/constructions/polish.lean @@ -46,7 +46,7 @@ analytic sets. -/ open set function polish_space pi_nat topological_space metric filter -open_locale topological_space measure_theory filter +open_locale topology measure_theory filter variables {α : Type*} [topological_space α] {ι : Type*} diff --git a/src/measure_theory/constructions/prod.lean b/src/measure_theory/constructions/prod.lean index 2f7ef21f9d5b7..43b1c74bd77e6 100644 --- a/src/measure_theory/constructions/prod.lean +++ b/src/measure_theory/constructions/prod.lean @@ -60,7 +60,7 @@ product measure, Fubini's theorem, Tonelli's theorem, Fubini-Tonelli theorem -/ noncomputable theory -open_locale classical topological_space ennreal measure_theory +open_locale classical topology ennreal measure_theory open set function real ennreal open measure_theory measurable_space measure_theory.measure open topological_space (hiding generate_from) diff --git a/src/measure_theory/covering/besicovitch.lean b/src/measure_theory/covering/besicovitch.lean index 186470555d8a2..c86b9f073a593 100644 --- a/src/measure_theory/covering/besicovitch.lean +++ b/src/measure_theory/covering/besicovitch.lean @@ -99,7 +99,7 @@ noncomputable theory universe u open metric set filter fin measure_theory topological_space -open_locale topological_space classical big_operators ennreal measure_theory nnreal +open_locale topology classical big_operators ennreal measure_theory nnreal /-! diff --git a/src/measure_theory/covering/besicovitch_vector_space.lean b/src/measure_theory/covering/besicovitch_vector_space.lean index 8c4b787747f5c..dc64f7093c509 100644 --- a/src/measure_theory/covering/besicovitch_vector_space.lean +++ b/src/measure_theory/covering/besicovitch_vector_space.lean @@ -43,7 +43,7 @@ In particular, this number is bounded by `5 ^ dim` by a straightforward measure universe u open metric set finite_dimensional measure_theory filter fin -open_locale ennreal topological_space +open_locale ennreal topology noncomputable theory diff --git a/src/measure_theory/covering/density_theorem.lean b/src/measure_theory/covering/density_theorem.lean index a17f343fa8511..0cc20e2988eb0 100644 --- a/src/measure_theory/covering/density_theorem.lean +++ b/src/measure_theory/covering/density_theorem.lean @@ -27,7 +27,7 @@ density theorem. noncomputable theory open set filter metric measure_theory topological_space -open_locale nnreal topological_space +open_locale nnreal topology namespace is_doubling_measure @@ -36,7 +36,7 @@ variables {α : Type*} [metric_space α] [measurable_space α] (μ : measure α) section variables [second_countable_topology α] [borel_space α] [is_locally_finite_measure μ] -open_locale topological_space +open_locale topology /-- A Vitali family in a space with a doubling measure, designed so that the sets at `x` contain all `closed_ball y r` when `dist x y ≤ K * r`. -/ diff --git a/src/measure_theory/covering/differentiation.lean b/src/measure_theory/covering/differentiation.lean index 1f9abd2f96386..ba4d4db70ca10 100644 --- a/src/measure_theory/covering/differentiation.lean +++ b/src/measure_theory/covering/differentiation.lean @@ -75,7 +75,7 @@ make no sense. However, the measure is not globally zero if the space is big eno -/ open measure_theory metric set filter topological_space measure_theory.measure -open_locale filter ennreal measure_theory nnreal topological_space +open_locale filter ennreal measure_theory nnreal topology variables {α : Type*} [metric_space α] {m0 : measurable_space α} {μ : measure α} (v : vitali_family μ) diff --git a/src/measure_theory/covering/liminf_limsup.lean b/src/measure_theory/covering/liminf_limsup.lean index 08303d9914277..911c1c887e9f7 100644 --- a/src/measure_theory/covering/liminf_limsup.lean +++ b/src/measure_theory/covering/liminf_limsup.lean @@ -24,7 +24,7 @@ carrying a doubling measure. -/ open set filter metric measure_theory topological_space -open_locale nnreal ennreal topological_space +open_locale nnreal ennreal topology variables {α : Type*} [metric_space α] [second_countable_topology α] [measurable_space α] [borel_space α] diff --git a/src/measure_theory/covering/one_dim.lean b/src/measure_theory/covering/one_dim.lean index 9476116be57c0..c1a8cb11bb377 100644 --- a/src/measure_theory/covering/one_dim.lean +++ b/src/measure_theory/covering/one_dim.lean @@ -15,7 +15,7 @@ by showing that intervals belong to the relevant Vitali family. -/ open set measure_theory is_doubling_measure filter -open_locale topological_space +open_locale topology namespace real diff --git a/src/measure_theory/covering/vitali.lean b/src/measure_theory/covering/vitali.lean index 11e767bed8b25..00142dc43e186 100644 --- a/src/measure_theory/covering/vitali.lean +++ b/src/measure_theory/covering/vitali.lean @@ -36,7 +36,7 @@ This version is given in `vitali.vitali_family`. variables {α ι : Type*} open set metric measure_theory topological_space filter -open_locale nnreal classical ennreal topological_space +open_locale nnreal classical ennreal topology namespace vitali diff --git a/src/measure_theory/covering/vitali_family.lean b/src/measure_theory/covering/vitali_family.lean index 2aad85dc9503d..a3b7d7b46cf2d 100644 --- a/src/measure_theory/covering/vitali_family.lean +++ b/src/measure_theory/covering/vitali_family.lean @@ -46,7 +46,7 @@ Vitali relations there) -/ open measure_theory metric set filter topological_space measure_theory.measure -open_locale filter measure_theory topological_space +open_locale filter measure_theory topology variables {α : Type*} [metric_space α] diff --git a/src/measure_theory/decomposition/unsigned_hahn.lean b/src/measure_theory/decomposition/unsigned_hahn.lean index e2d80b9594cf2..998240ec6148a 100644 --- a/src/measure_theory/decomposition/unsigned_hahn.lean +++ b/src/measure_theory/decomposition/unsigned_hahn.lean @@ -22,7 +22,7 @@ Hahn decomposition -/ open set filter -open_locale classical topological_space ennreal +open_locale classical topology ennreal namespace measure_theory diff --git a/src/measure_theory/function/ae_eq_fun.lean b/src/measure_theory/function/ae_eq_fun.lean index 9e969f4c8068e..a5a408f1d3d7d 100644 --- a/src/measure_theory/function/ae_eq_fun.lean +++ b/src/measure_theory/function/ae_eq_fun.lean @@ -68,7 +68,7 @@ function space, almost everywhere equal, `L⁰`, ae_eq_fun -/ noncomputable theory -open_locale classical ennreal topological_space +open_locale classical ennreal topology open set filter topological_space ennreal emetric measure_theory function variables {α β γ δ : Type*} [measurable_space α] {μ ν : measure α} diff --git a/src/measure_theory/function/ae_eq_of_integral.lean b/src/measure_theory/function/ae_eq_of_integral.lean index b847808f2a0ce..c83b0b93f2156 100644 --- a/src/measure_theory/function/ae_eq_of_integral.lean +++ b/src/measure_theory/function/ae_eq_of_integral.lean @@ -162,7 +162,7 @@ end section ennreal -open_locale topological_space +open_locale topology lemma ae_le_of_forall_set_lintegral_le_of_sigma_finite [sigma_finite μ] {f g : α → ℝ≥0∞} (hf : measurable f) (hg : measurable g) diff --git a/src/measure_theory/function/conditional_expectation/basic.lean b/src/measure_theory/function/conditional_expectation/basic.lean index 3576925d83dce..3e39ed9c4daae 100644 --- a/src/measure_theory/function/conditional_expectation/basic.lean +++ b/src/measure_theory/function/conditional_expectation/basic.lean @@ -76,7 +76,7 @@ conditional expectation, conditional expected value noncomputable theory open topological_space measure_theory.Lp filter continuous_linear_map -open_locale nnreal ennreal topological_space big_operators measure_theory +open_locale nnreal ennreal topology big_operators measure_theory namespace measure_theory diff --git a/src/measure_theory/function/conditional_expectation/indicator.lean b/src/measure_theory/function/conditional_expectation/indicator.lean index 75145fe2c871c..4925425dcf5c1 100644 --- a/src/measure_theory/function/conditional_expectation/indicator.lean +++ b/src/measure_theory/function/conditional_expectation/indicator.lean @@ -24,7 +24,7 @@ a restricted measure. noncomputable theory open topological_space measure_theory.Lp filter continuous_linear_map -open_locale nnreal ennreal topological_space big_operators measure_theory +open_locale nnreal ennreal topology big_operators measure_theory namespace measure_theory diff --git a/src/measure_theory/function/conditional_expectation/real.lean b/src/measure_theory/function/conditional_expectation/real.lean index a93d5a547b9e7..4fa0ccd6f1fd9 100644 --- a/src/measure_theory/function/conditional_expectation/real.lean +++ b/src/measure_theory/function/conditional_expectation/real.lean @@ -27,7 +27,7 @@ This file proves some results regarding the conditional expectation of real-valu noncomputable theory open topological_space measure_theory.Lp filter continuous_linear_map -open_locale nnreal ennreal topological_space big_operators measure_theory +open_locale nnreal ennreal topology big_operators measure_theory namespace measure_theory diff --git a/src/measure_theory/function/continuous_map_dense.lean b/src/measure_theory/function/continuous_map_dense.lean index 9db92100eb208..b01c24e6cec64 100644 --- a/src/measure_theory/function/continuous_map_dense.lean +++ b/src/measure_theory/function/continuous_map_dense.lean @@ -41,7 +41,7 @@ Vitali-Carathéodory theorem, in the file `measure_theory.vitali_caratheodory`. -/ -open_locale ennreal nnreal topological_space bounded_continuous_function +open_locale ennreal nnreal topology bounded_continuous_function open measure_theory topological_space continuous_map variables {α : Type*} [measurable_space α] [topological_space α] [normal_space α] [borel_space α] diff --git a/src/measure_theory/function/convergence_in_measure.lean b/src/measure_theory/function/convergence_in_measure.lean index e921708ec31e5..562f06d6eeb93 100644 --- a/src/measure_theory/function/convergence_in_measure.lean +++ b/src/measure_theory/function/convergence_in_measure.lean @@ -37,7 +37,7 @@ convergence in measure and other notions of convergence. -/ open topological_space filter -open_locale nnreal ennreal measure_theory topological_space +open_locale nnreal ennreal measure_theory topology namespace measure_theory diff --git a/src/measure_theory/function/egorov.lean b/src/measure_theory/function/egorov.lean index 54c9a7cc343ad..63bfb80eb9b2c 100644 --- a/src/measure_theory/function/egorov.lean +++ b/src/measure_theory/function/egorov.lean @@ -21,7 +21,7 @@ convergence in measure. -/ noncomputable theory -open_locale classical measure_theory nnreal ennreal topological_space +open_locale classical measure_theory nnreal ennreal topology namespace measure_theory diff --git a/src/measure_theory/function/jacobian.lean b/src/measure_theory/function/jacobian.lean index edb51c6fe39ff..1052d71d844a7 100644 --- a/src/measure_theory/function/jacobian.lean +++ b/src/measure_theory/function/jacobian.lean @@ -86,7 +86,7 @@ Change of variables in integrals open measure_theory measure_theory.measure metric filter set finite_dimensional asymptotics topological_space -open_locale nnreal ennreal topological_space pointwise +open_locale nnreal ennreal topology pointwise variables {E F : Type*} [normed_add_comm_group E] [normed_space ℝ E] [finite_dimensional ℝ E] [normed_add_comm_group F] [normed_space ℝ F] {s : set E} {f : E → E} {f' : E → E →L[ℝ] E} diff --git a/src/measure_theory/function/l1_space.lean b/src/measure_theory/function/l1_space.lean index 9caa52b58626c..7178483822149 100644 --- a/src/measure_theory/function/l1_space.lean +++ b/src/measure_theory/function/l1_space.lean @@ -47,7 +47,7 @@ integrable, function space, l1 noncomputable theory -open_locale classical topological_space big_operators ennreal measure_theory nnreal +open_locale classical topology big_operators ennreal measure_theory nnreal open set filter topological_space ennreal emetric measure_theory diff --git a/src/measure_theory/function/locally_integrable.lean b/src/measure_theory/function/locally_integrable.lean index fc98e48967e75..520eb5619f3dd 100644 --- a/src/measure_theory/function/locally_integrable.lean +++ b/src/measure_theory/function/locally_integrable.lean @@ -21,7 +21,7 @@ on compact sets. -/ open measure_theory measure_theory.measure set function topological_space -open_locale topological_space interval +open_locale topology interval variables {X Y E R : Type*} [measurable_space X] [topological_space X] variables [measurable_space Y] [topological_space Y] diff --git a/src/measure_theory/function/lp_space.lean b/src/measure_theory/function/lp_space.lean index 9c98bf1ea7d94..d9985fc0b00c7 100644 --- a/src/measure_theory/function/lp_space.lean +++ b/src/measure_theory/function/lp_space.lean @@ -76,7 +76,7 @@ function coercion from the coercion to almost everywhere defined functions. noncomputable theory open topological_space measure_theory filter -open_locale nnreal ennreal big_operators topological_space measure_theory +open_locale nnreal ennreal big_operators topology measure_theory variables {α E F G : Type*} {m m0 : measurable_space α} {p : ℝ≥0∞} {q : ℝ} {μ ν : measure α} [normed_add_comm_group E] [normed_add_comm_group F] [normed_add_comm_group G] diff --git a/src/measure_theory/function/simple_func_dense.lean b/src/measure_theory/function/simple_func_dense.lean index 524954f92c8c6..5c7ca36511d8f 100644 --- a/src/measure_theory/function/simple_func_dense.lean +++ b/src/measure_theory/function/simple_func_dense.lean @@ -31,7 +31,7 @@ by a sequence of simple functions. -/ open set function filter topological_space ennreal emetric finset -open_locale classical topological_space ennreal measure_theory big_operators +open_locale classical topology ennreal measure_theory big_operators variables {α β ι E F 𝕜 : Type*} noncomputable theory diff --git a/src/measure_theory/function/simple_func_dense_lp.lean b/src/measure_theory/function/simple_func_dense_lp.lean index 8eb99c82022c8..7185f7c68dd4b 100644 --- a/src/measure_theory/function/simple_func_dense_lp.lean +++ b/src/measure_theory/function/simple_func_dense_lp.lean @@ -41,7 +41,7 @@ For `E` finite-dimensional, simple functions `α →ₛ E` are dense in L^∞ -- noncomputable theory open set function filter topological_space ennreal emetric finset -open_locale classical topological_space ennreal measure_theory big_operators +open_locale classical topology ennreal measure_theory big_operators variables {α β ι E F 𝕜 : Type*} namespace measure_theory diff --git a/src/measure_theory/function/strongly_measurable/basic.lean b/src/measure_theory/function/strongly_measurable/basic.lean index be7925460e9af..d5414eb8b9c20 100644 --- a/src/measure_theory/function/strongly_measurable/basic.lean +++ b/src/measure_theory/function/strongly_measurable/basic.lean @@ -56,7 +56,7 @@ measurable functions, as a basis for the Bochner integral. -/ open measure_theory filter topological_space function set measure_theory.measure -open_locale ennreal topological_space measure_theory nnreal big_operators +open_locale ennreal topology measure_theory nnreal big_operators /-- The typeclass `second_countable_topology_either α β` registers the fact that at least one of the two spaces has second countable topology. This is the right assumption to ensure that continuous diff --git a/src/measure_theory/function/strongly_measurable/lp.lean b/src/measure_theory/function/strongly_measurable/lp.lean index 296d54884495b..d6cea24301e9e 100644 --- a/src/measure_theory/function/strongly_measurable/lp.lean +++ b/src/measure_theory/function/strongly_measurable/lp.lean @@ -25,7 +25,7 @@ Springer, 2016. -/ open measure_theory filter topological_space function -open_locale ennreal topological_space measure_theory +open_locale ennreal topology measure_theory namespace measure_theory diff --git a/src/measure_theory/function/uniform_integrable.lean b/src/measure_theory/function/uniform_integrable.lean index 789d0c95688e2..2add4cce0f1c7 100644 --- a/src/measure_theory/function/uniform_integrable.lean +++ b/src/measure_theory/function/uniform_integrable.lean @@ -43,7 +43,7 @@ uniform integrable, uniformly absolutely continuous integral, Vitali convergence -/ noncomputable theory -open_locale classical measure_theory nnreal ennreal topological_space big_operators +open_locale classical measure_theory nnreal ennreal topology big_operators namespace measure_theory diff --git a/src/measure_theory/group/action.lean b/src/measure_theory/group/action.lean index 250c467310882..3fa13676beccd 100644 --- a/src/measure_theory/group/action.lean +++ b/src/measure_theory/group/action.lean @@ -17,7 +17,7 @@ typeclass for measures invariant under action of an (additive or multiplicative) some basic properties of such measures. -/ -open_locale ennreal nnreal pointwise topological_space +open_locale ennreal nnreal pointwise topology open measure_theory measure_theory.measure set function namespace measure_theory diff --git a/src/measure_theory/group/add_circle.lean b/src/measure_theory/group/add_circle.lean index 93388f293ccf4..50e9ffcab8e9c 100644 --- a/src/measure_theory/group/add_circle.lean +++ b/src/measure_theory/group/add_circle.lean @@ -21,7 +21,7 @@ The file is a place to collect measure-theoretic results about the additive circ -/ open set function filter measure_theory measure_theory.measure metric -open_locale measure_theory pointwise big_operators topological_space ennreal +open_locale measure_theory pointwise big_operators topology ennreal namespace add_circle diff --git a/src/measure_theory/group/fundamental_domain.lean b/src/measure_theory/group/fundamental_domain.lean index 4d48c958b734d..69a214a0ac2fc 100644 --- a/src/measure_theory/group/fundamental_domain.lean +++ b/src/measure_theory/group/fundamental_domain.lean @@ -35,7 +35,7 @@ We also generate additive versions of all theorems in this file using the `to_ad Elements of `s` that do not belong to any other translate of `s`. -/ -open_locale ennreal pointwise topological_space nnreal ennreal measure_theory +open_locale ennreal pointwise topology nnreal ennreal measure_theory open measure_theory measure_theory.measure set function topological_space filter namespace measure_theory diff --git a/src/measure_theory/group/measure.lean b/src/measure_theory/group/measure.lean index b3d09a47c4246..b2e6534481b58 100644 --- a/src/measure_theory/group/measure.lean +++ b/src/measure_theory/group/measure.lean @@ -27,7 +27,7 @@ We also give analogues of all these notions in the additive world. noncomputable theory -open_locale nnreal ennreal pointwise big_operators topological_space +open_locale nnreal ennreal pointwise big_operators topology open has_inv set function measure_theory.measure filter variables {𝕜 G H : Type*} [measurable_space G] [measurable_space H] diff --git a/src/measure_theory/integral/average.lean b/src/measure_theory/integral/average.lean index 7ec0fab852c75..e11fa52b50090 100644 --- a/src/measure_theory/integral/average.lean +++ b/src/measure_theory/integral/average.lean @@ -28,7 +28,7 @@ integral, center mass, average value -/ open measure_theory measure_theory.measure metric set filter topological_space function -open_locale topological_space big_operators ennreal convex +open_locale topology big_operators ennreal convex variables {α E F : Type*} {m0 : measurable_space α} [normed_add_comm_group E] [normed_space ℝ E] [complete_space E] diff --git a/src/measure_theory/integral/bochner.lean b/src/measure_theory/integral/bochner.lean index 5c57984864bb2..5d898d490e1bd 100644 --- a/src/measure_theory/integral/bochner.lean +++ b/src/measure_theory/integral/bochner.lean @@ -140,7 +140,7 @@ Bochner integral, simple function, function space, Lebesgue dominated convergenc -/ noncomputable theory -open_locale topological_space big_operators nnreal ennreal measure_theory +open_locale topology big_operators nnreal ennreal measure_theory open set filter topological_space ennreal emetric namespace measure_theory diff --git a/src/measure_theory/integral/circle_integral.lean b/src/measure_theory/integral/circle_integral.lean index 72912cc375d0b..55b871ba566d9 100644 --- a/src/measure_theory/integral/circle_integral.lean +++ b/src/measure_theory/integral/circle_integral.lean @@ -69,7 +69,7 @@ variables {E : Type*} [normed_add_comm_group E] noncomputable theory -open_locale real nnreal interval pointwise topological_space +open_locale real nnreal interval pointwise topology open complex measure_theory topological_space metric function set filter asymptotics /-! diff --git a/src/measure_theory/integral/divergence_theorem.lean b/src/measure_theory/integral/divergence_theorem.lean index f458e7b18d655..cb324a0182956 100644 --- a/src/measure_theory/integral/divergence_theorem.lean +++ b/src/measure_theory/integral/divergence_theorem.lean @@ -46,7 +46,7 @@ divergence theorem, Bochner integral -/ open set finset topological_space function box_integral measure_theory filter -open_locale big_operators classical topological_space interval +open_locale big_operators classical topology interval universes u diff --git a/src/measure_theory/integral/integrable_on.lean b/src/measure_theory/integral/integrable_on.lean index 95432b32f04f5..1275358d89507 100644 --- a/src/measure_theory/integral/integrable_on.lean +++ b/src/measure_theory/integral/integrable_on.lean @@ -21,7 +21,7 @@ at `l`. noncomputable theory open set filter topological_space measure_theory function -open_locale classical topological_space interval big_operators filter ennreal measure_theory +open_locale classical topology interval big_operators filter ennreal measure_theory variables {α β E F : Type*} [measurable_space α] diff --git a/src/measure_theory/integral/integral_eq_improper.lean b/src/measure_theory/integral/integral_eq_improper.lean index 43b38176df3eb..5f583e481dd94 100644 --- a/src/measure_theory/integral/integral_eq_improper.lean +++ b/src/measure_theory/integral/integral_eq_improper.lean @@ -54,7 +54,7 @@ in analysis. -/ open measure_theory filter set topological_space -open_locale ennreal nnreal topological_space +open_locale ennreal nnreal topology namespace measure_theory diff --git a/src/measure_theory/integral/interval_integral.lean b/src/measure_theory/integral/interval_integral.lean index 8a99942ce3321..a2f650caf7302 100644 --- a/src/measure_theory/integral/interval_integral.lean +++ b/src/measure_theory/integral/interval_integral.lean @@ -166,7 +166,7 @@ noncomputable theory open topological_space (second_countable_topology) open measure_theory set classical filter function -open_locale classical topological_space filter ennreal big_operators interval nnreal +open_locale classical topology filter ennreal big_operators interval nnreal variables {ι 𝕜 E F A : Type*} [normed_add_comm_group E] diff --git a/src/measure_theory/integral/lebesgue.lean b/src/measure_theory/integral/lebesgue.lean index fa276733d378f..efb3bed88cd4f 100644 --- a/src/measure_theory/integral/lebesgue.lean +++ b/src/measure_theory/integral/lebesgue.lean @@ -36,7 +36,7 @@ We introduce the following notation for the lower Lebesgue integral of a functio noncomputable theory open set (hiding restrict restrict_apply) filter ennreal function (support) -open_locale classical topological_space big_operators nnreal ennreal measure_theory +open_locale classical topology big_operators nnreal ennreal measure_theory namespace measure_theory diff --git a/src/measure_theory/integral/set_integral.lean b/src/measure_theory/integral/set_integral.lean index 1ad13b8d3bbbc..6d8721454d123 100644 --- a/src/measure_theory/integral/set_integral.lean +++ b/src/measure_theory/integral/set_integral.lean @@ -48,7 +48,7 @@ but we reference them here because all theorems about set integrals are in this noncomputable theory open set filter topological_space measure_theory function -open_locale classical topological_space interval big_operators filter ennreal nnreal measure_theory +open_locale classical topology interval big_operators filter ennreal nnreal measure_theory variables {α β E F : Type*} [measurable_space α] diff --git a/src/measure_theory/integral/set_to_l1.lean b/src/measure_theory/integral/set_to_l1.lean index 97479a4bb2cfb..4b2fece04f2c5 100644 --- a/src/measure_theory/integral/set_to_l1.lean +++ b/src/measure_theory/integral/set_to_l1.lean @@ -68,7 +68,7 @@ with finite measure. Its value on other sets is ignored. -/ noncomputable theory -open_locale classical topological_space big_operators nnreal ennreal measure_theory pointwise +open_locale classical topology big_operators nnreal ennreal measure_theory pointwise open set filter topological_space ennreal emetric namespace measure_theory diff --git a/src/measure_theory/measure/doubling.lean b/src/measure_theory/measure/doubling.lean index c293a057d4dd6..2fbc89fc639ac 100644 --- a/src/measure_theory/measure/doubling.lean +++ b/src/measure_theory/measure/doubling.lean @@ -25,7 +25,7 @@ This file records basic files on doubling measures. noncomputable theory open set filter metric measure_theory topological_space -open_locale ennreal nnreal topological_space +open_locale ennreal nnreal topology /-- A measure `μ` is said to be a doubling measure if there exists a constant `C` such that for all sufficiently small radii `ε`, and for any centre, the measure of a ball of radius `2 * ε` is diff --git a/src/measure_theory/measure/finite_measure.lean b/src/measure_theory/measure/finite_measure.lean index 08a21dd2b591e..2f56da55396ca 100644 --- a/src/measure_theory/measure/finite_measure.lean +++ b/src/measure_theory/measure/finite_measure.lean @@ -73,7 +73,7 @@ open measure_theory open set open filter open bounded_continuous_function -open_locale topological_space ennreal nnreal bounded_continuous_function +open_locale topology ennreal nnreal bounded_continuous_function namespace measure_theory diff --git a/src/measure_theory/measure/haar.lean b/src/measure_theory/measure/haar.lean index 0883a2c3ba646..340eb715e8c07 100644 --- a/src/measure_theory/measure/haar.lean +++ b/src/measure_theory/measure/haar.lean @@ -62,7 +62,7 @@ where `ᵒ` denotes the interior. noncomputable theory open set has_inv function topological_space measurable_space -open_locale nnreal classical ennreal pointwise topological_space +open_locale nnreal classical ennreal pointwise topology namespace measure_theory namespace measure diff --git a/src/measure_theory/measure/haar_lebesgue.lean b/src/measure_theory/measure/haar_lebesgue.lean index b52c0c1122241..7cd63c3b50288 100644 --- a/src/measure_theory/measure/haar_lebesgue.lean +++ b/src/measure_theory/measure/haar_lebesgue.lean @@ -42,7 +42,7 @@ small `r`, see `eventually_nonempty_inter_smul_of_density_one`. -/ open topological_space set filter metric -open_locale ennreal pointwise topological_space nnreal +open_locale ennreal pointwise topology nnreal /-- The interval `[0,1]` as a compact set with non-empty interior. -/ def topological_space.positive_compacts.Icc01 : positive_compacts ℝ := diff --git a/src/measure_theory/measure/hausdorff.lean b/src/measure_theory/measure/hausdorff.lean index 0a7554a955e76..f25584df4b9b0 100644 --- a/src/measure_theory/measure/hausdorff.lean +++ b/src/measure_theory/measure/hausdorff.lean @@ -111,7 +111,7 @@ dimension. Hausdorff measure, measure, metric measure -/ -open_locale nnreal ennreal topological_space big_operators +open_locale nnreal ennreal topology big_operators open emetric set function filter encodable finite_dimensional topological_space diff --git a/src/measure_theory/measure/lebesgue.lean b/src/measure_theory/measure/lebesgue.lean index 8bd9fe374a079..2c8112475e211 100644 --- a/src/measure_theory/measure/lebesgue.lean +++ b/src/measure_theory/measure/lebesgue.lean @@ -29,7 +29,7 @@ are proved more generally for any additive Haar measure on a finite-dimensional noncomputable theory open classical set filter measure_theory measure_theory.measure topological_space open ennreal (of_real) -open_locale big_operators ennreal nnreal topological_space +open_locale big_operators ennreal nnreal topology /-! ### Definition of the Lebesgue measure and lengths of intervals diff --git a/src/measure_theory/measure/measure_space.lean b/src/measure_theory/measure/measure_space.lean index dc31c755e6e02..dca25b506f92e 100644 --- a/src/measure_theory/measure/measure_space.lean +++ b/src/measure_theory/measure/measure_space.lean @@ -90,7 +90,7 @@ measure, almost everywhere, measure space, completion, null set, null measurable noncomputable theory open set filter (hiding map) function measurable_space topological_space (second_countable_topology) -open_locale classical topological_space big_operators filter ennreal nnreal interval measure_theory +open_locale classical topology big_operators filter ennreal nnreal interval measure_theory variables {α β γ δ ι R R' : Type*} diff --git a/src/measure_theory/measure/measure_space_def.lean b/src/measure_theory/measure/measure_space_def.lean index 1f61e6e5d47b5..569a74f515b09 100644 --- a/src/measure_theory/measure/measure_space_def.lean +++ b/src/measure_theory/measure/measure_space_def.lean @@ -54,7 +54,7 @@ measure, almost everywhere, measure space noncomputable theory open classical set filter (hiding map) function measurable_space -open_locale classical topological_space big_operators filter ennreal nnreal +open_locale classical topology big_operators filter ennreal nnreal variables {α β γ δ ι : Type*} diff --git a/src/measure_theory/measure/open_pos.lean b/src/measure_theory/measure/open_pos.lean index acf862904523b..a0ef83afa07e0 100644 --- a/src/measure_theory/measure/open_pos.lean +++ b/src/measure_theory/measure/open_pos.lean @@ -15,7 +15,7 @@ about these measures. -/ -open_locale topological_space ennreal measure_theory +open_locale topology ennreal measure_theory open set function filter namespace measure_theory diff --git a/src/measure_theory/measure/outer_measure.lean b/src/measure_theory/measure/outer_measure.lean index 1c03e3594c176..5cdfc931f5781 100644 --- a/src/measure_theory/measure/outer_measure.lean +++ b/src/measure_theory/measure/outer_measure.lean @@ -54,7 +54,7 @@ outer measure, Carathéodory-measurable, Carathéodory's criterion noncomputable theory open set function filter topological_space (second_countable_topology) -open_locale classical big_operators nnreal topological_space ennreal measure_theory +open_locale classical big_operators nnreal topology ennreal measure_theory namespace measure_theory diff --git a/src/measure_theory/measure/portmanteau.lean b/src/measure_theory/measure/portmanteau.lean index 7bdf190a5e22d..debe09934a0d4 100644 --- a/src/measure_theory/measure/portmanteau.lean +++ b/src/measure_theory/measure/portmanteau.lean @@ -73,7 +73,7 @@ open measure_theory open set open filter open bounded_continuous_function -open_locale topological_space ennreal nnreal bounded_continuous_function +open_locale topology ennreal nnreal bounded_continuous_function namespace measure_theory diff --git a/src/measure_theory/measure/probability_measure.lean b/src/measure_theory/measure/probability_measure.lean index 0e0388c03b317..b56872b82a18a 100644 --- a/src/measure_theory/measure/probability_measure.lean +++ b/src/measure_theory/measure/probability_measure.lean @@ -69,7 +69,7 @@ open measure_theory open set open filter open bounded_continuous_function -open_locale topological_space ennreal nnreal bounded_continuous_function +open_locale topology ennreal nnreal bounded_continuous_function namespace measure_theory diff --git a/src/measure_theory/measure/regular.lean b/src/measure_theory/measure/regular.lean index 520260f52a41a..b5434ca92d550 100644 --- a/src/measure_theory/measure/regular.lean +++ b/src/measure_theory/measure/regular.lean @@ -128,7 +128,7 @@ proofs or statements do not apply directly. -/ open set filter -open_locale ennreal topological_space nnreal big_operators +open_locale ennreal topology nnreal big_operators namespace measure_theory namespace measure diff --git a/src/measure_theory/measure/stieltjes.lean b/src/measure_theory/measure/stieltjes.lean index da738a8787086..2c6a68198e461 100644 --- a/src/measure_theory/measure/stieltjes.lean +++ b/src/measure_theory/measure/stieltjes.lean @@ -25,7 +25,7 @@ a Borel measure `f.measure`. noncomputable theory open classical set filter function open ennreal (of_real) -open_locale big_operators ennreal nnreal topological_space measure_theory +open_locale big_operators ennreal nnreal topology measure_theory /-! ### Basic properties of Stieltjes functions -/ diff --git a/src/number_theory/liouville/liouville_with.lean b/src/number_theory/liouville/liouville_with.lean index 087dce9d8126e..c05c8aeaa9979 100644 --- a/src/number_theory/liouville/liouville_with.lean +++ b/src/number_theory/liouville/liouville_with.lean @@ -34,7 +34,7 @@ Liouville number, irrational, irrationality exponent -/ open filter metric real set -open_locale filter topological_space +open_locale filter topology /-- We say that a real number `x` is a Liouville number with exponent `p : ℝ` if there exists a real number `C` such that for infinitely many denominators `n` there exists a numerator `m` such that diff --git a/src/number_theory/liouville/measure.lean b/src/number_theory/liouville/measure.lean index 1b8f44c5241fc..c7527bdeb75e6 100644 --- a/src/number_theory/liouville/measure.lean +++ b/src/number_theory/liouville/measure.lean @@ -25,7 +25,7 @@ measure. The fact that the filters are disjoint means that two mutually exclusiv Liouville number, Lebesgue measure, residual, generic property -/ -open_locale filter big_operators ennreal topological_space nnreal +open_locale filter big_operators ennreal topology nnreal open filter set metric measure_theory real lemma set_of_liouville_with_subset_aux : diff --git a/src/number_theory/modular_forms/basic.lean b/src/number_theory/modular_forms/basic.lean index d2811244e051a..ba025ac6d38c5 100644 --- a/src/number_theory/modular_forms/basic.lean +++ b/src/number_theory/modular_forms/basic.lean @@ -20,7 +20,7 @@ modular form. open complex upper_half_plane -open_locale topological_space manifold upper_half_plane +open_locale topology manifold upper_half_plane noncomputable theory diff --git a/src/number_theory/padics/hensel.lean b/src/number_theory/padics/hensel.lean index 88d82eda6ebd2..57521f1b1a33d 100644 --- a/src/number_theory/padics/hensel.lean +++ b/src/number_theory/padics/hensel.lean @@ -33,7 +33,7 @@ p-adic, p adic, padic, p-adic integer noncomputable theory -open_locale classical topological_space +open_locale classical topology -- We begin with some general lemmas that are used below in the computation. diff --git a/src/number_theory/well_approximable.lean b/src/number_theory/well_approximable.lean index 621d66baf8be1..05e5a4fc84720 100644 --- a/src/number_theory/well_approximable.lean +++ b/src/number_theory/well_approximable.lean @@ -53,7 +53,7 @@ An elementary (non-measure-theoretic) argument shows that if `¬ hδ` holds then -/ open set filter function metric measure_theory -open_locale measure_theory topological_space pointwise +open_locale measure_theory topology pointwise /-- In a seminormed group `A`, given `n : ℕ` and `δ : ℝ`, `approx_order_of A n δ` is the set of elements within a distance `δ` of a point of order `n`. -/ diff --git a/src/order/filter/zero_and_bounded_at_filter.lean b/src/order/filter/zero_and_bounded_at_filter.lean index 49cdbe93dfba0..ef69104312417 100644 --- a/src/order/filter/zero_and_bounded_at_filter.lean +++ b/src/order/filter/zero_and_bounded_at_filter.lean @@ -21,7 +21,7 @@ namespace filter variables {α β : Type*} -open_locale topological_space +open_locale topology /-- If `l` is a filter on `α`, then a function `f : α → β` is `zero_at_filter l` if it tends to zero along `l`. -/ diff --git a/src/probability/borel_cantelli.lean b/src/probability/borel_cantelli.lean index 6951712b3d4e1..928753b73658e 100644 --- a/src/probability/borel_cantelli.lean +++ b/src/probability/borel_cantelli.lean @@ -22,7 +22,7 @@ filtration. -/ -open_locale measure_theory probability_theory ennreal big_operators topological_space +open_locale measure_theory probability_theory ennreal big_operators topology open measure_theory probability_theory measurable_space topological_space diff --git a/src/probability/ident_distrib.lean b/src/probability/ident_distrib.lean index 8a0ca1a2a7348..75775ebd5c98a 100644 --- a/src/probability/ident_distrib.lean +++ b/src/probability/ident_distrib.lean @@ -52,7 +52,7 @@ open measure_theory filter finset noncomputable theory -open_locale topological_space big_operators measure_theory ennreal nnreal +open_locale topology big_operators measure_theory ennreal nnreal variables {α β γ δ : Type*} [measurable_space α] [measurable_space β] [measurable_space γ] [measurable_space δ] diff --git a/src/probability/martingale/borel_cantelli.lean b/src/probability/martingale/borel_cantelli.lean index d5093ba542c98..363bfe0a98eee 100644 --- a/src/probability/martingale/borel_cantelli.lean +++ b/src/probability/martingale/borel_cantelli.lean @@ -28,7 +28,7 @@ is required to prove the generalized Borel-Cantelli. -/ open filter -open_locale nnreal ennreal measure_theory probability_theory big_operators topological_space +open_locale nnreal ennreal measure_theory probability_theory big_operators topology namespace measure_theory diff --git a/src/probability/martingale/convergence.lean b/src/probability/martingale/convergence.lean index 69fca0645ba62..f670120fad772 100644 --- a/src/probability/martingale/convergence.lean +++ b/src/probability/martingale/convergence.lean @@ -42,7 +42,7 @@ theorems. -/ open topological_space filter measure_theory.filtration -open_locale nnreal ennreal measure_theory probability_theory big_operators topological_space +open_locale nnreal ennreal measure_theory probability_theory big_operators topology namespace measure_theory diff --git a/src/probability/martingale/upcrossing.lean b/src/probability/martingale/upcrossing.lean index a333845c99b71..c2d3cb4769399 100644 --- a/src/probability/martingale/upcrossing.lean +++ b/src/probability/martingale/upcrossing.lean @@ -53,7 +53,7 @@ We mostly follow the proof from [Kallenberg, *Foundations of modern probability* -/ open topological_space filter -open_locale nnreal ennreal measure_theory probability_theory big_operators topological_space +open_locale nnreal ennreal measure_theory probability_theory big_operators topology namespace measure_theory diff --git a/src/probability/process/adapted.lean b/src/probability/process/adapted.lean index 21b1047306077..3cb162bc3b916 100644 --- a/src/probability/process/adapted.lean +++ b/src/probability/process/adapted.lean @@ -33,7 +33,7 @@ adapted, progressively measurable -/ open filter order topological_space -open_locale classical measure_theory nnreal ennreal topological_space big_operators +open_locale classical measure_theory nnreal ennreal topology big_operators namespace measure_theory diff --git a/src/probability/process/filtration.lean b/src/probability/process/filtration.lean index f498406b4836c..7a3c14f17260b 100644 --- a/src/probability/process/filtration.lean +++ b/src/probability/process/filtration.lean @@ -30,7 +30,7 @@ filtration, stochastic process -/ open filter order topological_space -open_locale classical measure_theory nnreal ennreal topological_space big_operators +open_locale classical measure_theory nnreal ennreal topology big_operators namespace measure_theory diff --git a/src/probability/process/hitting_time.lean b/src/probability/process/hitting_time.lean index 2e46740b5d738..7881a97d1152d 100644 --- a/src/probability/process/hitting_time.lean +++ b/src/probability/process/hitting_time.lean @@ -33,7 +33,7 @@ hitting times indexed by the natural numbers or the reals. By taking the bounds -/ open filter order topological_space -open_locale classical measure_theory nnreal ennreal topological_space big_operators +open_locale classical measure_theory nnreal ennreal topology big_operators namespace measure_theory diff --git a/src/probability/process/stopping.lean b/src/probability/process/stopping.lean index 424706ee2a303..29832258abf99 100644 --- a/src/probability/process/stopping.lean +++ b/src/probability/process/stopping.lean @@ -32,7 +32,7 @@ stopping time, stochastic process -/ open filter order topological_space -open_locale classical measure_theory nnreal ennreal topological_space big_operators +open_locale classical measure_theory nnreal ennreal topology big_operators namespace measure_theory diff --git a/src/probability/strong_law.lean b/src/probability/strong_law.lean index 4e2ca506d13aa..e09fa4393ca78 100644 --- a/src/probability/strong_law.lean +++ b/src/probability/strong_law.lean @@ -54,7 +54,7 @@ noncomputable theory open measure_theory filter finset asymptotics open set (indicator) -open_locale topological_space big_operators measure_theory probability_theory ennreal nnreal +open_locale topology big_operators measure_theory probability_theory ennreal nnreal namespace probability_theory diff --git a/src/topology/G_delta.lean b/src/topology/G_delta.lean index 71ae4d3de353e..f9f07550a2592 100644 --- a/src/topology/G_delta.lean +++ b/src/topology/G_delta.lean @@ -33,7 +33,7 @@ Gδ set, residual set -/ noncomputable theory -open_locale classical topological_space filter uniformity +open_locale classical topology filter uniformity open filter encodable set diff --git a/src/topology/alexandroff.lean b/src/topology/alexandroff.lean index d818326b6323d..2cc85d25521bb 100644 --- a/src/topology/alexandroff.lean +++ b/src/topology/alexandroff.lean @@ -34,7 +34,7 @@ one-point compactification, compactness -/ open set filter -open_locale classical topological_space filter +open_locale classical topology filter /-! ### Definition and basic properties diff --git a/src/topology/algebra/const_mul_action.lean b/src/topology/algebra/const_mul_action.lean index 6b253a88bfc34..af3be61886696 100644 --- a/src/topology/algebra/const_mul_action.lean +++ b/src/topology/algebra/const_mul_action.lean @@ -37,7 +37,7 @@ Hausdorff, discrete group, properly discontinuous, quotient space -/ -open_locale topological_space pointwise +open_locale topology pointwise open filter set topological_space diff --git a/src/topology/algebra/constructions.lean b/src/topology/algebra/constructions.lean index f2d7c3657c7c9..356f844454364 100644 --- a/src/topology/algebra/constructions.lean +++ b/src/topology/algebra/constructions.lean @@ -20,7 +20,7 @@ topological space, opposite monoid, units variables {M X : Type*} open filter -open_locale topological_space +open_locale topology namespace mul_opposite diff --git a/src/topology/algebra/field.lean b/src/topology/algebra/field.lean index 657ca380445ed..f46c2cd535940 100644 --- a/src/topology/algebra/field.lean +++ b/src/topology/algebra/field.lean @@ -178,7 +178,7 @@ end affine_homeomorph section local_extr variables {α β : Type*} [topological_space α] [linear_ordered_semifield β] {a : α} -open_locale topological_space +open_locale topology lemma is_local_min.inv {f : α → β} {a : α} (h1 : is_local_min f a) (h2 : ∀ᶠ z in 𝓝 a, 0 < f z) : is_local_max f⁻¹ a := diff --git a/src/topology/algebra/filter_basis.lean b/src/topology/algebra/filter_basis.lean index 5ef959246ca60..fe8f97746d553 100644 --- a/src/topology/algebra/filter_basis.lean +++ b/src/topology/algebra/filter_basis.lean @@ -35,7 +35,7 @@ Given a group `G` and a ring `R`: -/ open filter set topological_space function -open_locale topological_space filter pointwise +open_locale topology filter pointwise universe u diff --git a/src/topology/algebra/group/basic.lean b/src/topology/algebra/group/basic.lean index 7463aa2e243c2..e07a4c2f228eb 100644 --- a/src/topology/algebra/group/basic.lean +++ b/src/topology/algebra/group/basic.lean @@ -32,7 +32,7 @@ topological space, group, topological group -/ open classical set filter topological_space function -open_locale classical topological_space filter pointwise +open_locale classical topology filter pointwise universes u v w x variables {α : Type u} {β : Type v} {G : Type w} {H : Type x} diff --git a/src/topology/algebra/group/compact.lean b/src/topology/algebra/group/compact.lean index 0c223c744d154..62ce628e6263f 100644 --- a/src/topology/algebra/group/compact.lean +++ b/src/topology/algebra/group/compact.lean @@ -16,7 +16,7 @@ imports developing either positive compacts or the compact open topology. -/ open classical set filter topological_space function -open_locale classical topological_space filter pointwise +open_locale classical topology filter pointwise universes u v w x variables {α : Type u} {β : Type v} {G : Type w} {H : Type x} diff --git a/src/topology/algebra/group_with_zero.lean b/src/topology/algebra/group_with_zero.lean index 10f1d3ef63ad8..93b8fdccf51e8 100644 --- a/src/topology/algebra/group_with_zero.lean +++ b/src/topology/algebra/group_with_zero.lean @@ -29,7 +29,7 @@ On a `group_with_zero` with continuous multiplication, we also define left and r as homeomorphisms. -/ -open_locale topological_space filter +open_locale topology filter open filter function /-! diff --git a/src/topology/algebra/infinite_sum.lean b/src/topology/algebra/infinite_sum.lean index 7eaa1e5dc1354..c8a486f123e4d 100644 --- a/src/topology/algebra/infinite_sum.lean +++ b/src/topology/algebra/infinite_sum.lean @@ -28,7 +28,7 @@ generally, see `has_sum.tendsto_sum_nat`. noncomputable theory open finset filter function classical -open_locale topological_space classical big_operators nnreal +open_locale topology classical big_operators nnreal variables {α : Type*} {β : Type*} {γ : Type*} {δ : Type*} diff --git a/src/topology/algebra/module/basic.lean b/src/topology/algebra/module/basic.lean index e13655b155398..e3c2915ab023c 100644 --- a/src/topology/algebra/module/basic.lean +++ b/src/topology/algebra/module/basic.lean @@ -28,7 +28,7 @@ The corresponding notation for equivalences is `M ≃SL[σ] M₂`, `M ≃L[R] M -/ open filter linear_map (ker range) -open_locale topological_space big_operators filter +open_locale topology big_operators filter universes u v w u' diff --git a/src/topology/algebra/module/linear_pmap.lean b/src/topology/algebra/module/linear_pmap.lean index 2740693517f52..4580db88f328f 100644 --- a/src/topology/algebra/module/linear_pmap.lean +++ b/src/topology/algebra/module/linear_pmap.lean @@ -42,7 +42,7 @@ Unbounded operators, closed operators -/ -open_locale topological_space +open_locale topology variables {R E F : Type*} diff --git a/src/topology/algebra/module/locally_convex.lean b/src/topology/algebra/module/locally_convex.lean index 63a591c2b868e..11ea2c947006b 100644 --- a/src/topology/algebra/module/locally_convex.lean +++ b/src/topology/algebra/module/locally_convex.lean @@ -29,7 +29,7 @@ In a module, this is equivalent to `0` satisfying such properties. open topological_space filter set -open_locale topological_space pointwise +open_locale topology pointwise section semimodule diff --git a/src/topology/algebra/module/strong_topology.lean b/src/topology/algebra/module/strong_topology.lean index 9d2aeab58b238..880e183841291 100644 --- a/src/topology/algebra/module/strong_topology.lean +++ b/src/topology/algebra/module/strong_topology.lean @@ -55,7 +55,7 @@ sets). uniform convergence, bounded convergence -/ -open_locale topological_space uniform_convergence +open_locale topology uniform_convergence namespace continuous_linear_map diff --git a/src/topology/algebra/module/weak_dual.lean b/src/topology/algebra/module/weak_dual.lean index 5b1339fe4fb96..878a1b9f0b47a 100644 --- a/src/topology/algebra/module/weak_dual.lean +++ b/src/topology/algebra/module/weak_dual.lean @@ -61,7 +61,7 @@ weak-star, weak dual, duality noncomputable theory open filter -open_locale topological_space +open_locale topology variables {α 𝕜 𝕝 R E F M : Type*} diff --git a/src/topology/algebra/monoid.lean b/src/topology/algebra/monoid.lean index 1d4acc363497c..3cccb7d1c337b 100644 --- a/src/topology/algebra/monoid.lean +++ b/src/topology/algebra/monoid.lean @@ -18,7 +18,7 @@ the definitions. universes u v open classical set filter topological_space -open_locale classical topological_space big_operators pointwise +open_locale classical topology big_operators pointwise variables {ι α X M N : Type*} [topological_space X] diff --git a/src/topology/algebra/mul_action.lean b/src/topology/algebra/mul_action.lean index cfcc12f19a4d4..e0108ac9622d1 100644 --- a/src/topology/algebra/mul_action.lean +++ b/src/topology/algebra/mul_action.lean @@ -33,7 +33,7 @@ Besides homeomorphisms mentioned above, in this file we provide lemmas like `con or `filter.tendsto.smul` that provide dot-syntax access to `continuous_smul`. -/ -open_locale topological_space pointwise +open_locale topology pointwise open filter /-- Class `has_continuous_smul M X` says that the scalar multiplication `(•) : M → X → X` diff --git a/src/topology/algebra/nonarchimedean/adic_topology.lean b/src/topology/algebra/nonarchimedean/adic_topology.lean index 9bb1c72c97753..c5382d5f73c57 100644 --- a/src/topology/algebra/nonarchimedean/adic_topology.lean +++ b/src/topology/algebra/nonarchimedean/adic_topology.lean @@ -43,7 +43,7 @@ to make sure it is definitionally equal to the `I`-topology on `R` seen as a `R` variables {R : Type*} [comm_ring R] open set topological_add_group submodule filter -open_locale topological_space pointwise +open_locale topology pointwise namespace ideal diff --git a/src/topology/algebra/nonarchimedean/bases.lean b/src/topology/algebra/nonarchimedean/bases.lean index 2d43306ca2806..d1a63ae9a582b 100644 --- a/src/topology/algebra/nonarchimedean/bases.lean +++ b/src/topology/algebra/nonarchimedean/bases.lean @@ -28,7 +28,7 @@ sub-modules in a commutative algebra. This important example gives rises to the -/ open set filter function lattice add_group_with_zero_nhd -open_locale topological_space filter pointwise +open_locale topology filter pointwise /-- A family of additive subgroups on a ring `A` is a subgroups basis if it satisfies some axioms ensuring there is a topology on `A` which is compatible with the ring structure and diff --git a/src/topology/algebra/open_subgroup.lean b/src/topology/algebra/open_subgroup.lean index fd4eb90f9f208..d8d01b9f7022b 100644 --- a/src/topology/algebra/open_subgroup.lean +++ b/src/topology/algebra/open_subgroup.lean @@ -29,7 +29,7 @@ Note that this notion is especially relevant in a non-archimedean context, for i -/ open topological_space -open_locale topological_space +open_locale topology /-- The type of open subgroups of a topological additive group. -/ @[ancestor add_subgroup] diff --git a/src/topology/algebra/order/compact.lean b/src/topology/algebra/order/compact.lean index c2aa39290434e..153b9fe5e7bf6 100644 --- a/src/topology/algebra/order/compact.lean +++ b/src/topology/algebra/order/compact.lean @@ -25,7 +25,7 @@ compact, extreme value theorem -/ open filter order_dual topological_space function set -open_locale filter topological_space +open_locale filter topology /-! ### Compactness of a closed interval diff --git a/src/topology/algebra/order/extend_from.lean b/src/topology/algebra/order/extend_from.lean index 922ff11885cc5..41517fd80e2b5 100644 --- a/src/topology/algebra/order/extend_from.lean +++ b/src/topology/algebra/order/extend_from.lean @@ -11,7 +11,7 @@ import topology.extend_from -/ open filter set topological_space -open_locale topological_space classical +open_locale topology classical universes u v variables {α : Type u} {β : Type v} diff --git a/src/topology/algebra/order/extr_closure.lean b/src/topology/algebra/order/extr_closure.lean index 310426d2e8bda..0d4cdddf5e2b8 100644 --- a/src/topology/algebra/order/extr_closure.lean +++ b/src/topology/algebra/order/extr_closure.lean @@ -15,7 +15,7 @@ not) maximum (or minimum) on a set `s` at a point `a` and is continuous on the c -/ open filter set -open_locale topological_space +open_locale topology variables {X Y : Type*} [topological_space X] [topological_space Y] [preorder Y] [order_closed_topology Y] {f g : X → Y} {s : set X} {a : X} diff --git a/src/topology/algebra/order/field.lean b/src/topology/algebra/order/field.lean index 2b9eeecfab668..e6802847b5b39 100644 --- a/src/topology/algebra/order/field.lean +++ b/src/topology/algebra/order/field.lean @@ -17,7 +17,7 @@ import topology.algebra.field open set filter topological_space open function open order_dual (to_dual of_dual) -open_locale topological_space classical filter +open_locale topology classical filter variables {α β : Type*} variables [linear_ordered_field α] [topological_space α] [order_topology α] diff --git a/src/topology/algebra/order/filter.lean b/src/topology/algebra/order/filter.lean index 5ec5f68ec72c4..31ced314235d9 100644 --- a/src/topology/algebra/order/filter.lean +++ b/src/topology/algebra/order/filter.lean @@ -13,7 +13,7 @@ In this file we prove that `𝓝 (f x)` tends to `𝓝 filter.at_top` provided t `filter.at_top`, and similarly for `filter.at_bot`. -/ -open_locale topological_space +open_locale topology namespace filter diff --git a/src/topology/algebra/order/floor.lean b/src/topology/algebra/order/floor.lean index be5a80187cd7a..9b7ccd9ab235d 100644 --- a/src/topology/algebra/order/floor.lean +++ b/src/topology/algebra/order/floor.lean @@ -24,7 +24,7 @@ This file proves statements about limits and continuity of functions involving ` -/ open filter function int set -open_locale topological_space +open_locale topology variables {α β γ : Type*} [linear_ordered_ring α] [floor_ring α] diff --git a/src/topology/algebra/order/intermediate_value.lean b/src/topology/algebra/order/intermediate_value.lean index 77dad319a8919..e82eba6907532 100644 --- a/src/topology/algebra/order/intermediate_value.lean +++ b/src/topology/algebra/order/intermediate_value.lean @@ -37,7 +37,7 @@ intermediate value theorem, connected space, connected set -/ open filter order_dual topological_space function set -open_locale topological_space filter +open_locale topology filter universes u v w diff --git a/src/topology/algebra/order/left_right.lean b/src/topology/algebra/order/left_right.lean index 12a4d6582dd20..fe98ae82544fe 100644 --- a/src/topology/algebra/order/left_right.lean +++ b/src/topology/algebra/order/left_right.lean @@ -23,7 +23,7 @@ left continuous, right continuous -/ open set filter -open_locale topological_space +open_locale topology section partial_order diff --git a/src/topology/algebra/order/left_right_lim.lean b/src/topology/algebra/order/left_right_lim.lean index 688e2f22bd595..be928f32253b8 100644 --- a/src/topology/algebra/order/left_right_lim.lean +++ b/src/topology/algebra/order/left_right_lim.lean @@ -31,7 +31,7 @@ Prove corresponding stronger results for strict_mono and strict_anti functions. -/ open set filter -open_locale topological_space +open_locale topology section diff --git a/src/topology/algebra/order/liminf_limsup.lean b/src/topology/algebra/order/liminf_limsup.lean index 1ad1eb77d26bf..ef1bd9c837074 100644 --- a/src/topology/algebra/order/liminf_limsup.lean +++ b/src/topology/algebra/order/liminf_limsup.lean @@ -13,7 +13,7 @@ import topology.order.basic -/ open filter -open_locale topological_space classical +open_locale topology classical universes u v variables {α : Type u} {β : Type v} @@ -322,7 +322,7 @@ end monotone section infi_and_supr -open_locale topological_space +open_locale topology open filter set diff --git a/src/topology/algebra/order/monotone_continuity.lean b/src/topology/algebra/order/monotone_continuity.lean index d81ec0ab63b93..2b18d6f919796 100644 --- a/src/topology/algebra/order/monotone_continuity.lean +++ b/src/topology/algebra/order/monotone_continuity.lean @@ -21,7 +21,7 @@ continuous, monotone -/ open set filter -open_locale topological_space +open_locale topology section linear_order variables {α β : Type*} [linear_order α] [topological_space α] [order_topology α] diff --git a/src/topology/algebra/order/monotone_convergence.lean b/src/topology/algebra/order/monotone_convergence.lean index f2c2e1bfbb075..10724d464f353 100644 --- a/src/topology/algebra/order/monotone_convergence.lean +++ b/src/topology/algebra/order/monotone_convergence.lean @@ -26,7 +26,7 @@ monotone convergence -/ open filter set function -open_locale filter topological_space classical +open_locale filter topology classical variables {α β : Type*} diff --git a/src/topology/algebra/order/proj_Icc.lean b/src/topology/algebra/order/proj_Icc.lean index a9c3162be3515..c48e97502ec2a 100644 --- a/src/topology/algebra/order/proj_Icc.lean +++ b/src/topology/algebra/order/proj_Icc.lean @@ -14,7 +14,7 @@ to show that `Icc_extend h f` is continuous if and only if `f` is continuous. -/ open set filter -open_locale filter topological_space +open_locale filter topology variables {α β γ : Type*} [linear_order α] [topological_space γ] {a b c : α} {h : a ≤ b} diff --git a/src/topology/algebra/order/t5.lean b/src/topology/algebra/order/t5.lean index f2a48d583681c..2e3d851be9fcf 100644 --- a/src/topology/algebra/order/t5.lean +++ b/src/topology/algebra/order/t5.lean @@ -14,7 +14,7 @@ topological space. -/ open filter set function order_dual -open_locale topological_space filter interval +open_locale topology filter interval variables {X : Type*} [linear_order X] [topological_space X] [order_topology X] {a b c : X} {s t : set X} diff --git a/src/topology/algebra/ring.lean b/src/topology/algebra/ring.lean index 8b5d89f91bb39..899ca1d4bdd47 100644 --- a/src/topology/algebra/ring.lean +++ b/src/topology/algebra/ring.lean @@ -32,7 +32,7 @@ of topological (semi)rings. -/ open classical set filter topological_space function -open_locale classical topological_space filter +open_locale classical topology filter section topological_semiring variables (α : Type*) diff --git a/src/topology/algebra/star.lean b/src/topology/algebra/star.lean index adadf76af1dbd..99421351dbf6f 100644 --- a/src/topology/algebra/star.lean +++ b/src/topology/algebra/star.lean @@ -16,7 +16,7 @@ This file defines the `has_continuous_star` typeclass, along with instances on ` -/ -open_locale filter topological_space +open_locale filter topology open filter universes u diff --git a/src/topology/algebra/uniform_convergence.lean b/src/topology/algebra/uniform_convergence.lean index 796468b322c7f..fa23721dcb906 100644 --- a/src/topology/algebra/uniform_convergence.lean +++ b/src/topology/algebra/uniform_convergence.lean @@ -49,7 +49,7 @@ uniform convergence, strong dual -/ open filter -open_locale topological_space pointwise uniform_convergence +open_locale topology pointwise uniform_convergence section algebraic_instances diff --git a/src/topology/algebra/uniform_field.lean b/src/topology/algebra/uniform_field.lean index f8a53e5876f01..c1ac29242fc85 100644 --- a/src/topology/algebra/uniform_field.lean +++ b/src/topology/algebra/uniform_field.lean @@ -34,7 +34,7 @@ type class and the main results are the instances `uniform_space.completion.fiel noncomputable theory -open_locale classical uniformity topological_space +open_locale classical uniformity topology open set uniform_space uniform_space.completion filter diff --git a/src/topology/algebra/uniform_group.lean b/src/topology/algebra/uniform_group.lean index e731c5798fa64..d26baabfde3f8 100644 --- a/src/topology/algebra/uniform_group.lean +++ b/src/topology/algebra/uniform_group.lean @@ -34,7 +34,7 @@ group naturally induces a uniform structure. -/ noncomputable theory -open_locale classical uniformity topological_space filter pointwise +open_locale classical uniformity topology filter pointwise section uniform_group open filter set diff --git a/src/topology/algebra/valuation.lean b/src/topology/algebra/valuation.lean index 119ea43e16382..737093c2420b4 100644 --- a/src/topology/algebra/valuation.lean +++ b/src/topology/algebra/valuation.lean @@ -16,7 +16,7 @@ The main definition is a `valued` type class which equips a ring with a valuatio values in a group with zero. Other instances are then deduced from this. -/ -open_locale classical topological_space uniformity +open_locale classical topology uniformity open set valuation noncomputable theory diff --git a/src/topology/algebra/valued_field.lean b/src/topology/algebra/valued_field.lean index 469308e6f535e..849a9b8268ac1 100644 --- a/src/topology/algebra/valued_field.lean +++ b/src/topology/algebra/valued_field.lean @@ -29,7 +29,7 @@ Then we extend the valuation given on `K` to a valuation on `hat K`. -/ open filter set -open_locale topological_space +open_locale topology section division_ring diff --git a/src/topology/algebra/with_zero_topology.lean b/src/topology/algebra/with_zero_topology.lean index 7714790c9053b..16f04ede602c6 100644 --- a/src/topology/algebra/with_zero_topology.lean +++ b/src/topology/algebra/with_zero_topology.lean @@ -32,7 +32,7 @@ All other instances will (`ordered_topology`, `t3_space`, `has_continuous_mul`) -/ -open_locale topological_space filter +open_locale topology filter open topological_space filter set function namespace linear_ordered_comm_group_with_zero diff --git a/src/topology/bases.lean b/src/topology/bases.lean index 8771c8af835a9..29294bf7158e4 100644 --- a/src/topology/bases.lean +++ b/src/topology/bases.lean @@ -46,7 +46,7 @@ More fine grained instances for `first_countable_topology`, `separable_space`, ` -/ open set filter function -open_locale topological_space filter +open_locale topology filter noncomputable theory namespace topological_space diff --git a/src/topology/basic.lean b/src/topology/basic.lean index 5b88faadd7d39..743bed8e70c69 100644 --- a/src/topology/basic.lean +++ b/src/topology/basic.lean @@ -652,24 +652,18 @@ neighborhoods of `a` forms a filter, the neighborhood filter at `a`, is here def infimum over the principal filters of all open sets containing `a`. -/ @[irreducible] def nhds (a : α) : filter α := (⨅ s ∈ {s : set α | a ∈ s ∧ is_open s}, 𝓟 s) -localized "notation (name := nhds) `𝓝` := nhds" in topological_space +localized "notation (name := nhds) `𝓝` := nhds" in topology /-- The "neighborhood within" filter. Elements of `𝓝[s] a` are sets containing the intersection of `s` and a neighborhood of `a`. -/ def nhds_within (a : α) (s : set α) : filter α := 𝓝 a ⊓ 𝓟 s -localized "notation (name := nhds_within) - `𝓝[` s `] ` x:100 := nhds_within x s" in topological_space -localized "notation (name := nhds_within.ne) - `𝓝[≠] ` x:100 := nhds_within x {x}ᶜ" in topological_space -localized "notation (name := nhds_within.ge) - `𝓝[≥] ` x:100 := nhds_within x (set.Ici x)" in topological_space -localized "notation (name := nhds_within.le) - `𝓝[≤] ` x:100 := nhds_within x (set.Iic x)" in topological_space -localized "notation (name := nhds_within.gt) - `𝓝[>] ` x:100 := nhds_within x (set.Ioi x)" in topological_space -localized "notation (name := nhds_within.lt) - `𝓝[<] ` x:100 := nhds_within x (set.Iio x)" in topological_space +localized "notation (name := nhds_within) `𝓝[` s `] ` x:100 := nhds_within x s" in topology +localized "notation (name := nhds_within.ne) `𝓝[≠] ` x:100 := nhds_within x {x}ᶜ" in topology +localized "notation (name := nhds_within.ge) `𝓝[≥] ` x:100 := nhds_within x (set.Ici x)" in topology +localized "notation (name := nhds_within.le) `𝓝[≤] ` x:100 := nhds_within x (set.Iic x)" in topology +localized "notation (name := nhds_within.gt) `𝓝[>] ` x:100 := nhds_within x (set.Ioi x)" in topology +localized "notation (name := nhds_within.lt) `𝓝[<] ` x:100 := nhds_within x (set.Iio x)" in topology lemma nhds_def (a : α) : 𝓝 a = (⨅ s ∈ {s : set α | a ∈ s ∧ is_open s}, 𝓟 s) := by rw nhds @@ -1264,7 +1258,7 @@ end topological_space section continuous variables {α : Type*} {β : Type*} {γ : Type*} {δ : Type*} variables [topological_space α] [topological_space β] [topological_space γ] -open_locale topological_space +open_locale topology /-- A function between topological spaces is continuous if the preimage of every open set is open. Registered as a structure to make sure it is not unfolded by Lean. -/ diff --git a/src/topology/category/Compactum.lean b/src/topology/category/Compactum.lean index 1798f92eefc2a..6bc1fc6992906 100644 --- a/src/topology/category/Compactum.lean +++ b/src/topology/category/Compactum.lean @@ -72,7 +72,7 @@ We also add wrappers around structures which already exist. Here are the main on universe u open category_theory filter ultrafilter topological_space category_theory.limits has_finite_inter -open_locale classical topological_space +open_locale classical topology local notation `β` := of_type_monad ultrafilter diff --git a/src/topology/compact_open.lean b/src/topology/compact_open.lean index 2c41791bd4e02..4fd8e8b260178 100644 --- a/src/topology/compact_open.lean +++ b/src/topology/compact_open.lean @@ -34,7 +34,7 @@ compact-open, curry, function space -/ open set -open_locale topological_space +open_locale topology namespace continuous_map diff --git a/src/topology/connected.lean b/src/topology/connected.lean index 19602fb7b606b..269c05f25224d 100644 --- a/src/topology/connected.lean +++ b/src/topology/connected.lean @@ -40,7 +40,7 @@ https://ncatlab.org/nlab/show/too+simple+to+be+simple#relationship_to_biased_def -/ open set function topological_space relation -open_locale classical topological_space +open_locale classical topology universes u v variables {α : Type u} {β : Type v} {ι : Type*} {π : ι → Type*} [topological_space α] diff --git a/src/topology/constructions.lean b/src/topology/constructions.lean index 565bf50b7a7b7..ca83960d42f7a 100644 --- a/src/topology/constructions.lean +++ b/src/topology/constructions.lean @@ -34,7 +34,7 @@ product, sum, disjoint union, subspace, quotient space noncomputable theory open topological_space set filter function -open_locale classical topological_space filter +open_locale classical topology filter universes u v variables {α : Type u} {β : Type v} {γ δ ε ζ : Type*} diff --git a/src/topology/continuous_function/bounded.lean b/src/topology/continuous_function/bounded.lean index 7cbba6795e4d7..10b2f516edc71 100644 --- a/src/topology/continuous_function/bounded.lean +++ b/src/topology/continuous_function/bounded.lean @@ -19,7 +19,7 @@ the uniform distance. -/ noncomputable theory -open_locale topological_space classical nnreal uniformity uniform_convergence +open_locale topology classical nnreal uniformity uniform_convergence open set filter metric function diff --git a/src/topology/continuous_function/compact.lean b/src/topology/continuous_function/compact.lean index 191ccc57697a7..c20644015b441 100644 --- a/src/topology/continuous_function/compact.lean +++ b/src/topology/continuous_function/compact.lean @@ -24,7 +24,7 @@ you should restate it here. You can also use -/ noncomputable theory -open_locale topological_space classical nnreal bounded_continuous_function big_operators +open_locale topology classical nnreal bounded_continuous_function big_operators open set filter metric diff --git a/src/topology/continuous_function/stone_weierstrass.lean b/src/topology/continuous_function/stone_weierstrass.lean index 9756dccbdb3c3..b9de96ac298f1 100644 --- a/src/topology/continuous_function/stone_weierstrass.lean +++ b/src/topology/continuous_function/stone_weierstrass.lean @@ -160,7 +160,7 @@ begin exact h, end -open_locale topological_space +open_locale topology -- Here's the fun part of Stone-Weierstrass! theorem sublattice_closure_eq_top diff --git a/src/topology/continuous_function/zero_at_infty.lean b/src/topology/continuous_function/zero_at_infty.lean index 947c71539b09c..be667aa1b3767 100644 --- a/src/topology/continuous_function/zero_at_infty.lean +++ b/src/topology/continuous_function/zero_at_infty.lean @@ -24,7 +24,7 @@ universes u v w variables {F : Type*} {α : Type u} {β : Type v} {γ : Type w} [topological_space α] -open_locale bounded_continuous_function topological_space +open_locale bounded_continuous_function topology open filter metric /-- `C₀(α, β)` is the type of continuous functions `α → β` which vanish at infinity from a diff --git a/src/topology/continuous_on.lean b/src/topology/continuous_on.lean index 35dab90838e09..f2a69334ad418 100644 --- a/src/topology/continuous_on.lean +++ b/src/topology/continuous_on.lean @@ -27,7 +27,7 @@ equipped with the subspace topology. -/ open set filter function -open_locale topological_space filter +open_locale topology filter variables {α : Type*} {β : Type*} {γ : Type*} {δ : Type*} variables [topological_space α] diff --git a/src/topology/dense_embedding.lean b/src/topology/dense_embedding.lean index 261524bd0b5d8..b2ddb767158b8 100644 --- a/src/topology/dense_embedding.lean +++ b/src/topology/dense_embedding.lean @@ -25,7 +25,7 @@ has to be `dense_inducing` (not necessarily injective). noncomputable theory open set filter -open_locale classical topological_space filter +open_locale classical topology filter variables {α : Type*} {β : Type*} {γ : Type*} {δ : Type*} diff --git a/src/topology/extend_from.lean b/src/topology/extend_from.lean index 474a27784ae76..23838fd18eb1f 100644 --- a/src/topology/extend_from.lean +++ b/src/topology/extend_from.lean @@ -24,7 +24,7 @@ it suffices that `f` converges within `A` at any point of `B`, provided that noncomputable theory -open_locale topological_space +open_locale topology open filter set variables {X Y : Type*} [topological_space X] [topological_space Y] diff --git a/src/topology/fiber_bundle/basic.lean b/src/topology/fiber_bundle/basic.lean index 4eb51a3edfd8d..62cfece2a4958 100644 --- a/src/topology/fiber_bundle/basic.lean +++ b/src/topology/fiber_bundle/basic.lean @@ -167,7 +167,7 @@ Fiber bundle, topological bundle, structure group variables {ι : Type*} {B : Type*} {F : Type*} open topological_space filter set bundle -open_locale topological_space classical bundle +open_locale topology classical bundle attribute [mfld_simps] total_space.proj total_space_mk coe_fst coe_snd coe_snd_map_apply coe_snd_map_smul total_space.mk_cast diff --git a/src/topology/fiber_bundle/constructions.lean b/src/topology/fiber_bundle/constructions.lean index 4b3df9b72c90e..5c51d6cf9d543 100644 --- a/src/topology/fiber_bundle/constructions.lean +++ b/src/topology/fiber_bundle/constructions.lean @@ -25,7 +25,7 @@ fiber bundle, fibre bundle, fiberwise product, pullback -/ open topological_space filter set bundle -open_locale topological_space classical bundle +open_locale topology classical bundle /-! ### The trivial bundle -/ diff --git a/src/topology/fiber_bundle/trivialization.lean b/src/topology/fiber_bundle/trivialization.lean index 5efdcb77912b0..8e3a2ac2cdfd5 100644 --- a/src/topology/fiber_bundle/trivialization.lean +++ b/src/topology/fiber_bundle/trivialization.lean @@ -48,7 +48,7 @@ type of linear trivializations is not even particularly well-behaved. -/ open topological_space filter set bundle -open_locale topological_space classical bundle +open_locale topology classical bundle variables {ι : Type*} {B : Type*} {F : Type*} {E : B → Type*} variables (F) {Z : Type*} [topological_space B] [topological_space F] {proj : Z → B} diff --git a/src/topology/filter.lean b/src/topology/filter.lean index dd0d531b3fa43..fa87fc9e8dd6e 100644 --- a/src/topology/filter.lean +++ b/src/topology/filter.lean @@ -32,7 +32,7 @@ filter, topological space -/ open set filter topological_space -open_locale filter topological_space +open_locale filter topology variables {ι : Sort*} {α β X Y : Type*} diff --git a/src/topology/homeomorph.lean b/src/topology/homeomorph.lean index 3d0ef7172be6c..9997b1fca2832 100644 --- a/src/topology/homeomorph.lean +++ b/src/topology/homeomorph.lean @@ -27,7 +27,7 @@ directions continuous. We denote homeomorphisms with the notation `≃ₜ`. -/ open set filter -open_locale topological_space +open_locale topology variables {α : Type*} {β : Type*} {γ : Type*} {δ : Type*} diff --git a/src/topology/homotopy/homotopy_group.lean b/src/topology/homotopy/homotopy_group.lean index 929002c56e848..045f1d84ad6d6 100644 --- a/src/topology/homotopy/homotopy_group.lean +++ b/src/topology/homotopy/homotopy_group.lean @@ -29,7 +29,7 @@ TODO: show that `π n x` is a group when `n > 0` and abelian when `n > 1`. Show -/ -open_locale unit_interval topological_space +open_locale unit_interval topology noncomputable theory diff --git a/src/topology/inseparable.lean b/src/topology/inseparable.lean index 48f4bc2623521..f5dd2759c0953 100644 --- a/src/topology/inseparable.lean +++ b/src/topology/inseparable.lean @@ -35,7 +35,7 @@ topological space, separation setoid -/ open set filter function -open_locale topological_space filter +open_locale topology filter variables {X Y Z α ι : Type*} {π : ι → Type*} [topological_space X] [topological_space Y] [topological_space Z] [∀ i, topological_space (π i)] {x y z : X} {s : set X} {f : X → Y} diff --git a/src/topology/instances/add_circle.lean b/src/topology/instances/add_circle.lean index 371831392e71f..4e626f5d7c06e 100644 --- a/src/topology/instances/add_circle.lean +++ b/src/topology/instances/add_circle.lean @@ -50,7 +50,7 @@ the rational circle `add_circle (1 : ℚ)`, and so we set things up more general noncomputable theory open set function add_subgroup topological_space -open_locale topological_space +open_locale topology variables {𝕜 B : Type*} diff --git a/src/topology/instances/ennreal.lean b/src/topology/instances/ennreal.lean index 5bf8b42b77410..35b53f56688e4 100644 --- a/src/topology/instances/ennreal.lean +++ b/src/topology/instances/ennreal.lean @@ -13,7 +13,7 @@ import analysis.normed.group.basic noncomputable theory open classical set filter metric -open_locale classical topological_space ennreal nnreal big_operators filter +open_locale classical topology ennreal nnreal big_operators filter variables {α : Type*} {β : Type*} {γ : Type*} diff --git a/src/topology/instances/ereal.lean b/src/topology/instances/ereal.lean index 9b013eef3cba4..39238619680bc 100644 --- a/src/topology/instances/ereal.lean +++ b/src/topology/instances/ereal.lean @@ -28,7 +28,7 @@ Most proofs are adapted from the corresponding proofs on `ℝ≥0∞`. noncomputable theory open classical set filter metric topological_space -open_locale classical topological_space ennreal nnreal big_operators filter +open_locale classical topology ennreal nnreal big_operators filter variables {α : Type*} [topological_space α] diff --git a/src/topology/instances/irrational.lean b/src/topology/instances/irrational.lean index 1fe1ed9145659..357cfcadc2810 100644 --- a/src/topology/instances/irrational.lean +++ b/src/topology/instances/irrational.lean @@ -29,7 +29,7 @@ irrational, residual -/ open set filter metric -open_locale filter topological_space +open_locale filter topology lemma is_Gδ_irrational : is_Gδ {x | irrational x} := (countable_range _).is_Gδ_compl diff --git a/src/topology/instances/nnreal.lean b/src/topology/instances/nnreal.lean index 99e88c45e7175..2837afba8385a 100644 --- a/src/topology/instances/nnreal.lean +++ b/src/topology/instances/nnreal.lean @@ -45,7 +45,7 @@ a few of which rely on the fact that subtraction is continuous. -/ noncomputable theory open set topological_space metric filter -open_locale topological_space +open_locale topology namespace nnreal open_locale nnreal big_operators filter diff --git a/src/topology/instances/rat_lemmas.lean b/src/topology/instances/rat_lemmas.lean index b8e6c0c7e4373..641939a78f8ba 100644 --- a/src/topology/instances/rat_lemmas.lean +++ b/src/topology/instances/rat_lemmas.lean @@ -27,7 +27,7 @@ compactification. -/ open set metric filter topological_space -open_locale topological_space alexandroff +open_locale topology alexandroff local notation `ℚ∞` := alexandroff ℚ namespace rat diff --git a/src/topology/instances/real.lean b/src/topology/instances/real.lean index 28d5af70a62b1..65780cfc86e31 100644 --- a/src/topology/instances/real.lean +++ b/src/topology/instances/real.lean @@ -21,7 +21,7 @@ import topology.instances.int noncomputable theory open classical filter int metric set topological_space -open_locale classical topological_space filter uniformity interval +open_locale classical topology filter uniformity interval universes u v w variables {α : Type u} {β : Type v} {γ : Type w} diff --git a/src/topology/is_locally_homeomorph.lean b/src/topology/is_locally_homeomorph.lean index aa7d201b6050c..cdd16b8b411b7 100644 --- a/src/topology/is_locally_homeomorph.lean +++ b/src/topology/is_locally_homeomorph.lean @@ -20,7 +20,7 @@ This file defines local homeomorphisms. `local_homeomorph`, which is a homeomorphism between specific open subsets. -/ -open_locale topological_space +open_locale topology variables {X Y Z : Type*} [topological_space X] [topological_space Y] [topological_space Z] (g : Y → Z) (f : X → Y) (s : set X) (t : set Y) diff --git a/src/topology/list.lean b/src/topology/list.lean index 1569325904374..a7a761a61516b 100644 --- a/src/topology/list.lean +++ b/src/topology/list.lean @@ -10,7 +10,7 @@ import topology.algebra.monoid -/ open topological_space set filter -open_locale topological_space filter +open_locale topology filter variables {α : Type*} {β : Type*} [topological_space α] [topological_space β] diff --git a/src/topology/local_at_target.lean b/src/topology/local_at_target.lean index 5c258bb4eddc3..9a32dd7651a53 100644 --- a/src/topology/local_at_target.lean +++ b/src/topology/local_at_target.lean @@ -17,7 +17,7 @@ We show that the following properties of continuous maps are local at the target -/ open topological_space set filter -open_locale topological_space filter +open_locale topology filter variables {α β : Type*} [topological_space α] [topological_space β] {f : α → β} variables {s : set β} {ι : Type*} {U : ι → opens β} (hU : supr U = ⊤) diff --git a/src/topology/local_extr.lean b/src/topology/local_extr.lean index ed73d7aeb1c31..9ed89b5e3a917 100644 --- a/src/topology/local_extr.lean +++ b/src/topology/local_extr.lean @@ -36,7 +36,7 @@ universes u v w x variables {α : Type u} {β : Type v} {γ : Type w} {δ : Type x} [topological_space α] open set filter -open_locale topological_space filter +open_locale topology filter section preorder diff --git a/src/topology/local_homeomorph.lean b/src/topology/local_homeomorph.lean index a362dcc1b96f5..21de0e24902e4 100644 --- a/src/topology/local_homeomorph.lean +++ b/src/topology/local_homeomorph.lean @@ -43,7 +43,7 @@ then it should use `e.source ∩ s` or `e.target ∩ t`, not `s ∩ e.source` or -/ open function set filter topological_space (second_countable_topology) -open_locale topological_space +open_locale topology variables {α : Type*} {β : Type*} {γ : Type*} {δ : Type*} [topological_space α] [topological_space β] [topological_space γ] [topological_space δ] diff --git a/src/topology/locally_constant/basic.lean b/src/topology/locally_constant/basic.lean index ffae47b95d1ed..84cbb39b4d66c 100644 --- a/src/topology/locally_constant/basic.lean +++ b/src/topology/locally_constant/basic.lean @@ -28,7 +28,7 @@ This file sets up the theory of locally constant function from a topological spa variables {X Y Z α : Type*} [topological_space X] open set filter -open_locale topological_space +open_locale topology /-- A function between topological spaces is locally constant if the preimage of any set is open. -/ def is_locally_constant (f : X → Y) : Prop := ∀ s : set Y, is_open (f ⁻¹' s) diff --git a/src/topology/locally_finite.lean b/src/topology/locally_finite.lean index 50127489a3924..21bcd74e204d8 100644 --- a/src/topology/locally_finite.lean +++ b/src/topology/locally_finite.lean @@ -18,7 +18,7 @@ In this file we give the definition and prove basic properties of locally finite /- locally finite family [General Topology (Bourbaki, 1995)] -/ open set function filter -open_locale topological_space filter +open_locale topology filter universe u variables {ι : Type u} {ι' α X Y : Type*} [topological_space X] [topological_space Y] diff --git a/src/topology/maps.lean b/src/topology/maps.lean index 458748676cac0..38cc9151f2ec3 100644 --- a/src/topology/maps.lean +++ b/src/topology/maps.lean @@ -43,7 +43,7 @@ open map, closed map, embedding, quotient map, identification map -/ open set filter function -open_locale topological_space filter +open_locale topology filter variables {α : Type*} {β : Type*} {γ : Type*} {δ : Type*} diff --git a/src/topology/metric_space/baire.lean b/src/topology/metric_space/baire.lean index e76428dec6f66..7de9a219c1aec 100644 --- a/src/topology/metric_space/baire.lean +++ b/src/topology/metric_space/baire.lean @@ -25,7 +25,7 @@ has the countable intersection property. noncomputable theory -open_locale classical topological_space filter ennreal +open_locale classical topology filter ennreal open filter encodable set topological_space diff --git a/src/topology/metric_space/basic.lean b/src/topology/metric_space/basic.lean index 33b8f2bdc72d5..68d8dbc87ee33 100644 --- a/src/topology/metric_space/basic.lean +++ b/src/topology/metric_space/basic.lean @@ -49,7 +49,7 @@ metric, pseudo_metric, dist open set filter topological_space bornology -open_locale uniformity topological_space big_operators filter nnreal ennreal +open_locale uniformity topology big_operators filter nnreal ennreal universes u v w variables {α : Type u} {β : Type v} {X ι : Type*} diff --git a/src/topology/metric_space/cau_seq_filter.lean b/src/topology/metric_space/cau_seq_filter.lean index ac6401bc72230..11f3c9b94fe6d 100644 --- a/src/topology/metric_space/cau_seq_filter.lean +++ b/src/topology/metric_space/cau_seq_filter.lean @@ -15,7 +15,7 @@ of `cau_seq` Cauchy sequences. universes u v open set filter -open_locale topological_space classical +open_locale topology classical variable {β : Type v} diff --git a/src/topology/metric_space/closeds.lean b/src/topology/metric_space/closeds.lean index e08a06bb4a4bd..12e5998f38e63 100644 --- a/src/topology/metric_space/closeds.lean +++ b/src/topology/metric_space/closeds.lean @@ -23,7 +23,7 @@ always finite in this context. -/ noncomputable theory -open_locale classical topological_space ennreal +open_locale classical topology ennreal universe u open classical set function topological_space filter diff --git a/src/topology/metric_space/completion.lean b/src/topology/metric_space/completion.lean index 29c31422613e8..566a0eff7aa80 100644 --- a/src/topology/metric_space/completion.lean +++ b/src/topology/metric_space/completion.lean @@ -17,7 +17,7 @@ it defines the same uniformity as the already defined uniform structure on the c -/ open set filter uniform_space metric -open_locale filter topological_space uniformity +open_locale filter topology uniformity noncomputable theory universes u v diff --git a/src/topology/metric_space/contracting.lean b/src/topology/metric_space/contracting.lean index 8daeb4f8b0b6e..743f5fe0e4d63 100644 --- a/src/topology/metric_space/contracting.lean +++ b/src/topology/metric_space/contracting.lean @@ -27,7 +27,7 @@ of convergence, and some properties of the map sending a contracting map to its contracting map, fixed point, Banach fixed point theorem -/ -open_locale nnreal topological_space classical ennreal +open_locale nnreal topology classical ennreal open filter function variables {α : Type*} diff --git a/src/topology/metric_space/emetric_paracompact.lean b/src/topology/metric_space/emetric_paracompact.lean index 0d46182da5e39..e12583452acaf 100644 --- a/src/topology/metric_space/emetric_paracompact.lean +++ b/src/topology/metric_space/emetric_paracompact.lean @@ -23,7 +23,7 @@ metric space, paracompact space, normal space variable {α : Type*} -open_locale ennreal topological_space +open_locale ennreal topology open set namespace emetric diff --git a/src/topology/metric_space/emetric_space.lean b/src/topology/metric_space/emetric_space.lean index 1146f5df46403..cbbd0a29d972f 100644 --- a/src/topology/metric_space/emetric_space.lean +++ b/src/topology/metric_space/emetric_space.lean @@ -29,7 +29,7 @@ to `emetric_space` at the end. open set filter classical -open_locale uniformity topological_space big_operators filter nnreal ennreal +open_locale uniformity topology big_operators filter nnreal ennreal universes u v w variables {α : Type u} {β : Type v} {X : Type*} diff --git a/src/topology/metric_space/equicontinuity.lean b/src/topology/metric_space/equicontinuity.lean index f0b61162702b8..e8e2225e49443 100644 --- a/src/topology/metric_space/equicontinuity.lean +++ b/src/topology/metric_space/equicontinuity.lean @@ -32,7 +32,7 @@ equicontinuity, continuity modulus -/ open filter -open_locale topological_space uniformity +open_locale topology uniformity variables {α β ι : Type*} [pseudo_metric_space α] diff --git a/src/topology/metric_space/gluing.lean b/src/topology/metric_space/gluing.lean index 15e54523b73a8..6c628b30260c5 100644 --- a/src/topology/metric_space/gluing.lean +++ b/src/topology/metric_space/gluing.lean @@ -482,7 +482,7 @@ end local attribute [instance] sigma.metric_space -open_locale topological_space +open_locale topology open filter /-- The injection of a space in a disjoint union is an isometry -/ diff --git a/src/topology/metric_space/gromov_hausdorff.lean b/src/topology/metric_space/gromov_hausdorff.lean index 19158f8e22df5..ace9c9a3d64e9 100644 --- a/src/topology/metric_space/gromov_hausdorff.lean +++ b/src/topology/metric_space/gromov_hausdorff.lean @@ -40,7 +40,7 @@ i.e., it is complete and second countable. We also prove the Gromov compactness -/ noncomputable theory -open_locale classical topological_space ennreal +open_locale classical topology ennreal local notation `ℓ_infty_ℝ`:= lp (λ n : ℕ, ℝ) ∞ diff --git a/src/topology/metric_space/gromov_hausdorff_realized.lean b/src/topology/metric_space/gromov_hausdorff_realized.lean index c8bf97b84234c..f53cd689243bf 100644 --- a/src/topology/metric_space/gromov_hausdorff_realized.lean +++ b/src/topology/metric_space/gromov_hausdorff_realized.lean @@ -30,7 +30,7 @@ space structure on `X ⊕ Y`. The corresponding metric quotient is `optimal_GH_c -/ noncomputable theory -open_locale classical topological_space nnreal +open_locale classical topology nnreal universes u v w open classical set function topological_space filter metric quotient diff --git a/src/topology/metric_space/hausdorff_dimension.lean b/src/topology/metric_space/hausdorff_dimension.lean index a40f4865e133c..9267ca31287b9 100644 --- a/src/topology/metric_space/hausdorff_dimension.lean +++ b/src/topology/metric_space/hausdorff_dimension.lean @@ -79,7 +79,7 @@ We use the following notation localized in `measure_theory`. It is defined in Hausdorff measure, Hausdorff dimension, dimension -/ -open_locale measure_theory ennreal nnreal topological_space +open_locale measure_theory ennreal nnreal topology open measure_theory measure_theory.measure set topological_space finite_dimensional filter variables {ι X Y : Type*} [emetric_space X] [emetric_space Y] diff --git a/src/topology/metric_space/hausdorff_distance.lean b/src/topology/metric_space/hausdorff_distance.lean index 768bf966e24d7..d2c2a35abd69d 100644 --- a/src/topology/metric_space/hausdorff_distance.lean +++ b/src/topology/metric_space/hausdorff_distance.lean @@ -27,7 +27,7 @@ This files introduces: * `cthickening δ s`, the closed thickening by radius `δ` of a set `s` in a pseudo emetric space. -/ noncomputable theory -open_locale classical nnreal ennreal topological_space +open_locale classical nnreal ennreal topology universes u v w open classical set function topological_space filter diff --git a/src/topology/metric_space/holder.lean b/src/topology/metric_space/holder.lean index 6a4e9cb72c388..2db96800cd5eb 100644 --- a/src/topology/metric_space/holder.lean +++ b/src/topology/metric_space/holder.lean @@ -35,7 +35,7 @@ Hölder continuity, Lipschitz continuity variables {X Y Z : Type*} open filter set -open_locale nnreal ennreal topological_space +open_locale nnreal ennreal topology section emetric diff --git a/src/topology/metric_space/isometry.lean b/src/topology/metric_space/isometry.lean index 32b3f5e4bd66d..a54e0143dbcd5 100644 --- a/src/topology/metric_space/isometry.lean +++ b/src/topology/metric_space/isometry.lean @@ -23,7 +23,7 @@ universes u v w variables {α : Type u} {β : Type v} {γ : Type w} open function set -open_locale topological_space ennreal +open_locale topology ennreal /-- An isometry (also known as isometric embedding) is a map preserving the edistance between pseudoemetric spaces, or equivalently the distance between pseudometric space. -/ diff --git a/src/topology/metric_space/lipschitz.lean b/src/topology/metric_space/lipschitz.lean index 4f33443074d1e..331ca41e42b1a 100644 --- a/src/topology/metric_space/lipschitz.lean +++ b/src/topology/metric_space/lipschitz.lean @@ -40,7 +40,7 @@ argument, and return `lipschitz_with (real.to_nnreal K) f`. universes u v w x open filter function set -open_locale topological_space nnreal ennreal +open_locale topology nnreal ennreal variables {α : Type u} {β : Type v} {γ : Type w} {ι : Type x} diff --git a/src/topology/metric_space/metrizable.lean b/src/topology/metric_space/metrizable.lean index 250b0a3467d31..7fb64fc6a9b69 100644 --- a/src/topology/metric_space/metrizable.lean +++ b/src/topology/metric_space/metrizable.lean @@ -21,7 +21,7 @@ space structure. -/ open set filter metric -open_locale bounded_continuous_function filter topological_space +open_locale bounded_continuous_function filter topology namespace topological_space diff --git a/src/topology/metric_space/partition_of_unity.lean b/src/topology/metric_space/partition_of_unity.lean index 22f30453a6e2b..9b79f18b41305 100644 --- a/src/topology/metric_space/partition_of_unity.lean +++ b/src/topology/metric_space/partition_of_unity.lean @@ -24,7 +24,7 @@ lemma. metric space, partition of unity, locally finite -/ -open_locale topological_space ennreal big_operators nnreal filter +open_locale topology ennreal big_operators nnreal filter open set function filter topological_space variables {ι X : Type*} diff --git a/src/topology/metric_space/pi_nat.lean b/src/topology/metric_space/pi_nat.lean index ba35f8bb48a1a..1038852859a4c 100644 --- a/src/topology/metric_space/pi_nat.lean +++ b/src/topology/metric_space/pi_nat.lean @@ -49,7 +49,7 @@ in general), and `ι` is countable. -/ noncomputable theory -open_locale classical topological_space filter +open_locale classical topology filter open topological_space set metric filter function local attribute [simp] pow_le_pow_iff one_lt_two inv_le_inv @@ -801,7 +801,7 @@ lemma dist_le_dist_pi_of_dist_lt {x y : Π i, F i} {i : ι} (h : dist x y < (1/2 dist (x i) (y i) ≤ dist x y := by simpa only [not_le.2 h, false_or] using min_le_iff.1 (min_dist_le_dist_pi x y i) -open_locale big_operators topological_space +open_locale big_operators topology open filter open_locale nnreal diff --git a/src/topology/metric_space/polish.lean b/src/topology/metric_space/polish.lean index 058a8153b6f0d..253f5d1a46f89 100644 --- a/src/topology/metric_space/polish.lean +++ b/src/topology/metric_space/polish.lean @@ -44,7 +44,7 @@ with additional properties: -/ noncomputable theory -open_locale classical topological_space filter +open_locale classical topology filter open topological_space set metric filter function variables {α : Type*} {β : Type*} diff --git a/src/topology/metric_space/shrinking_lemma.lean b/src/topology/metric_space/shrinking_lemma.lean index 4833ac1741977..18f1c38122218 100644 --- a/src/topology/metric_space/shrinking_lemma.lean +++ b/src/topology/metric_space/shrinking_lemma.lean @@ -20,7 +20,7 @@ shrinking lemma, metric space universes u v open set metric -open_locale topological_space +open_locale topology variables {α : Type u} {ι : Type v} [metric_space α] [proper_space α] {c : ι → α} variables {x : α} {r : ℝ} {s : set α} diff --git a/src/topology/metric_space/thickened_indicator.lean b/src/topology/metric_space/thickened_indicator.lean index 59f690929ae74..71d0829959d5b 100644 --- a/src/topology/metric_space/thickened_indicator.lean +++ b/src/topology/metric_space/thickened_indicator.lean @@ -33,7 +33,7 @@ members of the approximating sequence are nonnegative bounded continuous functio -/ noncomputable theory -open_locale classical nnreal ennreal topological_space bounded_continuous_function +open_locale classical nnreal ennreal topology bounded_continuous_function open nnreal ennreal set metric emetric filter diff --git a/src/topology/nhds_set.lean b/src/topology/nhds_set.lean index a4a7a1916354e..9f8adbbdb44aa 100644 --- a/src/topology/nhds_set.lean +++ b/src/topology/nhds_set.lean @@ -24,7 +24,7 @@ Furthermore, we have the following results: -/ open set filter -open_locale topological_space filter +open_locale topology filter variables {α β : Type*} [topological_space α] [topological_space β] {s t s₁ s₂ t₁ t₂ : set α} {x : α} @@ -33,7 +33,7 @@ variables {α β : Type*} [topological_space α] [topological_space β] def nhds_set (s : set α) : filter α := Sup (nhds '' s) -localized "notation (name := nhds_set) `𝓝ˢ` := nhds_set" in topological_space +localized "notation (name := nhds_set) `𝓝ˢ` := nhds_set" in topology lemma nhds_set_diagonal (α) [topological_space (α × α)] : 𝓝ˢ (diagonal α) = ⨆ x, 𝓝 (x, x) := by { rw [nhds_set, ← range_diag, ← range_comp], refl } diff --git a/src/topology/order.lean b/src/topology/order.lean index c3c56da713d69..03acaaacd2b62 100644 --- a/src/topology/order.lean +++ b/src/topology/order.lean @@ -42,7 +42,7 @@ finer, coarser, induced topology, coinduced topology -/ open function set filter -open_locale classical topological_space filter +open_locale classical topology filter universes u v w diff --git a/src/topology/order/basic.lean b/src/topology/order/basic.lean index 3e2512eff625d..be90cee93ee77 100644 --- a/src/topology/order/basic.lean +++ b/src/topology/order/basic.lean @@ -77,7 +77,7 @@ for specific types. open set filter topological_space open function open order_dual (to_dual of_dual) -open_locale topological_space classical filter +open_locale topology classical filter universes u v w variables {α : Type u} {β : Type v} {γ : Type w} diff --git a/src/topology/order/lattice.lean b/src/topology/order/lattice.lean index 9cf03943e0905..6817698750018 100644 --- a/src/topology/order/lattice.lean +++ b/src/topology/order/lattice.lean @@ -23,7 +23,7 @@ topological, lattice -/ open filter -open_locale topological_space +open_locale topology /-- Let `L` be a topological space and let `L×L` be equipped with the product topology and let diff --git a/src/topology/paracompact.lean b/src/topology/paracompact.lean index 41204626271ad..505621d65b45e 100644 --- a/src/topology/paracompact.lean +++ b/src/topology/paracompact.lean @@ -46,7 +46,7 @@ compact space, paracompact space, locally finite covering -/ open set filter function -open_locale filter topological_space +open_locale filter topology universes u v diff --git a/src/topology/partial.lean b/src/topology/partial.lean index c09f9b41ab63f..4ebcbae24f64e 100644 --- a/src/topology/partial.lean +++ b/src/topology/partial.lean @@ -14,7 +14,7 @@ In this file we prove properties of `filter.ptendsto` etc in topological spaces. -/ open filter -open_locale topological_space +open_locale topology variables {α β : Type*} [topological_space α] diff --git a/src/topology/partition_of_unity.lean b/src/topology/partition_of_unity.lean index 638019acbeeed..866967b080840 100644 --- a/src/topology/partition_of_unity.lean +++ b/src/topology/partition_of_unity.lean @@ -75,7 +75,7 @@ partition of unity, bump function, Urysohn's lemma, normal space, paracompact sp universes u v open function set filter -open_locale big_operators topological_space classical +open_locale big_operators topology classical noncomputable theory diff --git a/src/topology/path_connected.lean b/src/topology/path_connected.lean index 33be284ce494c..76e5e9542fd13 100644 --- a/src/topology/path_connected.lean +++ b/src/topology/path_connected.lean @@ -60,7 +60,7 @@ on `(-∞, 0]` and to `y` on `[1, +∞)`. -/ noncomputable theory -open_locale classical topological_space filter unit_interval +open_locale classical topology filter unit_interval open filter set function unit_interval variables {X Y : Type*} [topological_space X] [topological_space Y] {x y z : X} {ι : Type*} diff --git a/src/topology/perfect.lean b/src/topology/perfect.lean index f788259dc54ae..051687e4c4dba 100644 --- a/src/topology/perfect.lean +++ b/src/topology/perfect.lean @@ -44,7 +44,7 @@ accumulation point, perfect set, Cantor-Bendixson. -/ -open_locale topological_space filter +open_locale topology filter open topological_space filter set variables {α : Type*} [topological_space α] {C : set α} diff --git a/src/topology/semicontinuous.lean b/src/topology/semicontinuous.lean index 1a42cd958b721..7d8121174b2e1 100644 --- a/src/topology/semicontinuous.lean +++ b/src/topology/semicontinuous.lean @@ -51,7 +51,7 @@ ones for lower semicontinuous functions using `order_dual`. -/ -open_locale topological_space big_operators ennreal +open_locale topology big_operators ennreal open set function filter variables {α : Type*} [topological_space α] {β : Type*} [preorder β] diff --git a/src/topology/separation.lean b/src/topology/separation.lean index 3e55ac79f6534..a1f19d14f4381 100644 --- a/src/topology/separation.lean +++ b/src/topology/separation.lean @@ -85,7 +85,7 @@ https://en.wikipedia.org/wiki/Separation_axiom -/ open function set filter topological_space -open_locale topological_space filter classical +open_locale topology filter classical universes u v variables {α : Type u} {β : Type v} [topological_space α] diff --git a/src/topology/sequences.lean b/src/topology/sequences.lean index 5b71cc2becf91..c5c15d868ba3c 100644 --- a/src/topology/sequences.lean +++ b/src/topology/sequences.lean @@ -58,7 +58,7 @@ sequentially closed, sequentially compact, sequential space -/ open set function filter topological_space -open_locale topological_space filter +open_locale topology filter variables {X Y : Type*} diff --git a/src/topology/stone_cech.lean b/src/topology/stone_cech.lean index 8f8340bd77ad6..b5410a75a9467 100644 --- a/src/topology/stone_cech.lean +++ b/src/topology/stone_cech.lean @@ -17,7 +17,7 @@ by Marius Stekelenburg, particularly section 5. noncomputable theory open filter set -open_locale topological_space +open_locale topology universes u v diff --git a/src/topology/subset_properties.lean b/src/topology/subset_properties.lean index 205223f8efb92..c255c89a188b4 100644 --- a/src/topology/subset_properties.lean +++ b/src/topology/subset_properties.lean @@ -49,7 +49,7 @@ https://ncatlab.org/nlab/show/too+simple+to+be+simple#relationship_to_biased_def -/ open set filter classical topological_space -open_locale classical topological_space filter +open_locale classical topology filter universes u v variables {α : Type u} {β : Type v} {ι : Type*} {π : ι → Type*} diff --git a/src/topology/support.lean b/src/topology/support.lean index 4beaa91af3bba..5e312403ab023 100644 --- a/src/topology/support.lean +++ b/src/topology/support.lean @@ -28,7 +28,7 @@ Furthermore, we say that `f` has compact support if the topological support of ` -/ open function set filter -open_locale topological_space +open_locale topology variables {X α α' β γ δ M E R : Type*} diff --git a/src/topology/tietze_extension.lean b/src/topology/tietze_extension.lean index abaea016c04ac..2e43880370bce 100644 --- a/src/topology/tietze_extension.lean +++ b/src/topology/tietze_extension.lean @@ -35,7 +35,7 @@ Tietze extension theorem, Urysohn's lemma, normal topological space variables {X Y : Type*} [topological_space X] [topological_space Y] [normal_space Y] open metric set filter -open_locale bounded_continuous_function topological_space +open_locale bounded_continuous_function topology noncomputable theory namespace bounded_continuous_function diff --git a/src/topology/uniform_space/basic.lean b/src/topology/uniform_space/basic.lean index 6e08e4ec9d8e3..aa67e33aabecb 100644 --- a/src/topology/uniform_space/basic.lean +++ b/src/topology/uniform_space/basic.lean @@ -112,7 +112,7 @@ But it makes a more systematic use of the filter library. -/ open set filter classical -open_locale classical topological_space filter +open_locale classical topology filter set_option eqn_compiler.zeta true diff --git a/src/topology/uniform_space/cauchy.lean b/src/topology/uniform_space/cauchy.lean index f648dcc997395..5425a039488cb 100644 --- a/src/topology/uniform_space/cauchy.lean +++ b/src/topology/uniform_space/cauchy.lean @@ -11,7 +11,7 @@ import topology.uniform_space.basic universes u v open filter topological_space set classical uniform_space function -open_locale classical uniformity topological_space filter +open_locale classical uniformity topology filter variables {α : Type u} {β : Type v} [uniform_space α] diff --git a/src/topology/uniform_space/compact.lean b/src/topology/uniform_space/compact.lean index 75744f7b46199..bb21c542896f3 100644 --- a/src/topology/uniform_space/compact.lean +++ b/src/topology/uniform_space/compact.lean @@ -32,7 +32,7 @@ loop. uniform space, uniform continuity, compact space -/ -open_locale classical uniformity topological_space filter +open_locale classical uniformity topology filter open filter uniform_space set variables {α β γ : Type*} [uniform_space α] [uniform_space β] diff --git a/src/topology/uniform_space/compact_convergence.lean b/src/topology/uniform_space/compact_convergence.lean index 7b3fc75c6ad8b..33f714537b1a1 100644 --- a/src/topology/uniform_space/compact_convergence.lean +++ b/src/topology/uniform_space/compact_convergence.lean @@ -76,7 +76,7 @@ of the uniform space structure on `C(α, β)` definitionally equal to the compac universes u₁ u₂ u₃ -open_locale filter uniformity topological_space +open_locale filter uniformity topology open uniform_space set filter variables {α : Type u₁} {β : Type u₂} [topological_space α] [uniform_space β] diff --git a/src/topology/uniform_space/complete_separated.lean b/src/topology/uniform_space/complete_separated.lean index 47c05bccd5af3..da9939d6a9848 100644 --- a/src/topology/uniform_space/complete_separated.lean +++ b/src/topology/uniform_space/complete_separated.lean @@ -14,7 +14,7 @@ This file is for elementary lemmas that depend on both Cauchy filters and separa -/ open filter -open_locale topological_space filter +open_locale topology filter variables {α : Type*} diff --git a/src/topology/uniform_space/completion.lean b/src/topology/uniform_space/completion.lean index 2158ae9f3b529..f954d08426b0b 100644 --- a/src/topology/uniform_space/completion.lean +++ b/src/topology/uniform_space/completion.lean @@ -44,7 +44,7 @@ noncomputable theory open filter set universes u v w x -open_locale uniformity classical topological_space filter +open_locale uniformity classical topology filter /-- Space of Cauchy filters diff --git a/src/topology/uniform_space/equicontinuity.lean b/src/topology/uniform_space/equicontinuity.lean index f8fa1f4166975..5a71e9c33d6db 100644 --- a/src/topology/uniform_space/equicontinuity.lean +++ b/src/topology/uniform_space/equicontinuity.lean @@ -81,7 +81,7 @@ equicontinuity, uniform convergence, ascoli section open uniform_space filter set -open_locale uniformity topological_space uniform_convergence +open_locale uniformity topology uniform_convergence variables {ι κ X Y Z α β γ 𝓕 : Type*} [topological_space X] [topological_space Y] [topological_space Z] [uniform_space α] [uniform_space β] [uniform_space γ] diff --git a/src/topology/uniform_space/matrix.lean b/src/topology/uniform_space/matrix.lean index a39badbdaf63a..2ad9ecbce6da8 100644 --- a/src/topology/uniform_space/matrix.lean +++ b/src/topology/uniform_space/matrix.lean @@ -10,7 +10,7 @@ import data.matrix.basic # Uniform space structure on matrices -/ -open_locale uniformity topological_space +open_locale uniformity topology variables (m n 𝕜 : Type*) [uniform_space 𝕜] diff --git a/src/topology/uniform_space/pi.lean b/src/topology/uniform_space/pi.lean index 7b0bedb535794..12ca6fde6a6a9 100644 --- a/src/topology/uniform_space/pi.lean +++ b/src/topology/uniform_space/pi.lean @@ -12,7 +12,7 @@ import topology.uniform_space.separation noncomputable theory -open_locale uniformity topological_space +open_locale uniformity topology section open filter uniform_space diff --git a/src/topology/uniform_space/separation.lean b/src/topology/uniform_space/separation.lean index 39c1ea89295cc..1c016b1d08084 100644 --- a/src/topology/uniform_space/separation.lean +++ b/src/topology/uniform_space/separation.lean @@ -66,7 +66,7 @@ uniformly continuous). -/ open filter topological_space set classical function uniform_space -open_locale classical topological_space uniformity filter +open_locale classical topology uniformity filter noncomputable theory set_option eqn_compiler.zeta true diff --git a/src/topology/uniform_space/uniform_convergence.lean b/src/topology/uniform_space/uniform_convergence.lean index 1ee5495cedb46..a1e146c7de013 100644 --- a/src/topology/uniform_space/uniform_convergence.lean +++ b/src/topology/uniform_space/uniform_convergence.lean @@ -60,7 +60,7 @@ Uniform limit, uniform convergence, tends uniformly to -/ noncomputable theory -open_locale topological_space classical uniformity filter +open_locale topology classical uniformity filter open set filter diff --git a/src/topology/uniform_space/uniform_convergence_topology.lean b/src/topology/uniform_space/uniform_convergence_topology.lean index 7b4610f323f91..51e5a7e1eeac8 100644 --- a/src/topology/uniform_space/uniform_convergence_topology.lean +++ b/src/topology/uniform_space/uniform_convergence_topology.lean @@ -131,7 +131,7 @@ uniform convergence -/ noncomputable theory -open_locale topological_space classical uniformity filter +open_locale topology classical uniformity filter open set filter diff --git a/src/topology/uniform_space/uniform_embedding.lean b/src/topology/uniform_space/uniform_embedding.lean index c38a4deeafc8a..b6958a215268f 100644 --- a/src/topology/uniform_space/uniform_embedding.lean +++ b/src/topology/uniform_space/uniform_embedding.lean @@ -14,7 +14,7 @@ Extension of uniform continuous functions. -/ open filter topological_space set classical -open_locale classical uniformity topological_space filter +open_locale classical uniformity topology filter section variables {α : Type*} {β : Type*} {γ : Type*} diff --git a/src/topology/unit_interval.lean b/src/topology/unit_interval.lean index 098dfaf050564..49f5f358b8253 100644 --- a/src/topology/unit_interval.lean +++ b/src/topology/unit_interval.lean @@ -19,7 +19,7 @@ We provide basic instances, as well as a custom tactic for discharging -/ noncomputable theory -open_locale classical topological_space filter +open_locale classical topology filter open set int set.Icc /-! ### The unit interval -/ diff --git a/src/topology/urysohns_lemma.lean b/src/topology/urysohns_lemma.lean index 32ea00924cfce..e9236a07a162f 100644 --- a/src/topology/urysohns_lemma.lean +++ b/src/topology/urysohns_lemma.lean @@ -71,7 +71,7 @@ Urysohn's lemma, normal topological space variables {X : Type*} [topological_space X] open set filter topological_space -open_locale topological_space filter +open_locale topology filter namespace urysohns From fac369018417f980cec5fcdafc766a69f88d8cfe Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Sat, 28 Jan 2023 07:02:44 +0000 Subject: [PATCH 0380/1291] chore(*): add mathlib4 synchronization comments (#18319) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.graded_monoid` * `algebra.module.equiv` * `algebra.module.opposites` * `algebra.module.submodule.basic` * `data.dfinsupp.basic` * `data.enat.lattice` * `data.finsupp.defs` * `data.finsupp.order` * `data.int.interval` * `data.list.cycle` * `data.nat.hyperoperation` * `data.sym.sym2` * `deprecated.subgroup` * `dynamics.periodic_pts` * `group_theory.free_group` * `group_theory.group_action.conj_act` * `group_theory.group_action.fixing_subgroup` * `group_theory.group_action.sub_mul_action` * `group_theory.is_free_group` * `group_theory.perm.list` * `group_theory.subgroup.saturated` * `group_theory.subgroup.simple` * `ring_theory.ring_invo` * `topology.bornology.hom` * `topology.locally_finite` * `topology.nhds_set` --- src/algebra/graded_monoid.lean | 3 +++ src/algebra/module/equiv.lean | 3 +++ src/algebra/module/opposites.lean | 3 +++ src/algebra/module/submodule/basic.lean | 3 +++ src/data/dfinsupp/basic.lean | 3 +++ src/data/enat/lattice.lean | 3 +++ src/data/finsupp/defs.lean | 3 +++ src/data/finsupp/order.lean | 3 +++ src/data/int/interval.lean | 3 +++ src/data/list/cycle.lean | 3 +++ src/data/nat/hyperoperation.lean | 3 +++ src/data/sym/sym2.lean | 3 +++ src/deprecated/subgroup.lean | 3 +++ src/dynamics/periodic_pts.lean | 3 +++ src/group_theory/free_group.lean | 3 +++ src/group_theory/group_action/conj_act.lean | 3 +++ src/group_theory/group_action/fixing_subgroup.lean | 3 +++ src/group_theory/group_action/sub_mul_action.lean | 3 +++ src/group_theory/is_free_group.lean | 3 +++ src/group_theory/perm/list.lean | 3 +++ src/group_theory/subgroup/saturated.lean | 3 +++ src/group_theory/subgroup/simple.lean | 3 +++ src/ring_theory/ring_invo.lean | 3 +++ src/topology/bornology/hom.lean | 3 +++ src/topology/locally_finite.lean | 3 +++ src/topology/nhds_set.lean | 3 +++ 26 files changed, 78 insertions(+) diff --git a/src/algebra/graded_monoid.lean b/src/algebra/graded_monoid.lean index a0c0af93134a6..454603c0a8341 100644 --- a/src/algebra/graded_monoid.lean +++ b/src/algebra/graded_monoid.lean @@ -14,6 +14,9 @@ import data.sigma.basic /-! # Additively-graded multiplicative structures +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This module provides a set of heterogeneous typeclasses for defining a multiplicative structure over the sigma type `graded_monoid A` such that `(*) : A i → A j → A (i + j)`; that is to say, `A` forms an additively-graded monoid. The typeclasses are: diff --git a/src/algebra/module/equiv.lean b/src/algebra/module/equiv.lean index a9fc1ee36fbae..636be037e8be8 100644 --- a/src/algebra/module/equiv.lean +++ b/src/algebra/module/equiv.lean @@ -9,6 +9,9 @@ import algebra.module.linear_map /-! # (Semi)linear equivalences +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define * `linear_equiv σ M M₂`, `M ≃ₛₗ[σ] M₂`: an invertible semilinear map. Here, `σ` is a `ring_hom` diff --git a/src/algebra/module/opposites.lean b/src/algebra/module/opposites.lean index f6f6a9ad0fcc2..d66b4be1d6cbd 100644 --- a/src/algebra/module/opposites.lean +++ b/src/algebra/module/opposites.lean @@ -9,6 +9,9 @@ import group_theory.group_action.opposite /-! # Module operations on `Mᵐᵒᵖ` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains definitions that build on top of the group action definitions in `group_theory.group_action.opposite`. -/ diff --git a/src/algebra/module/submodule/basic.lean b/src/algebra/module/submodule/basic.lean index 18cc10c629849..7e9547166f1ad 100644 --- a/src/algebra/module/submodule/basic.lean +++ b/src/algebra/module/submodule/basic.lean @@ -12,6 +12,9 @@ import group_theory.submonoid.membership # Submodules of a module +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define * `submodule R M` : a subset of a `module` `M` that contains zero and is closed with respect to diff --git a/src/data/dfinsupp/basic.lean b/src/data/dfinsupp/basic.lean index 8d649239ef552..f5327b314b0bb 100644 --- a/src/data/dfinsupp/basic.lean +++ b/src/data/dfinsupp/basic.lean @@ -13,6 +13,9 @@ import data.finset.preimage /-! # Dependent functions with finite support +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + For a non-dependent version see `data/finsupp.lean`. ## Notation diff --git a/src/data/enat/lattice.lean b/src/data/enat/lattice.lean index 73da7e8364477..887e933c7a73c 100644 --- a/src/data/enat/lattice.lean +++ b/src/data/enat/lattice.lean @@ -9,6 +9,9 @@ import data.enat.basic /-! # Extended natural numbers form a complete linear order +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This instance is not in `data.enat.basic` to avoid dependency on `finset`s. -/ diff --git a/src/data/finsupp/defs.lean b/src/data/finsupp/defs.lean index 4ba86d3d97737..e8623391f3093 100644 --- a/src/data/finsupp/defs.lean +++ b/src/data/finsupp/defs.lean @@ -9,6 +9,9 @@ import group_theory.submonoid.basic /-! # Type of functions with finite support +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + For any type `α` and any type `M` with zero, we define the type `finsupp α M` (notation: `α →₀ M`) of finitely supported functions from `α` to `M`, i.e. the functions which are zero everywhere on `α` except on a finite set. diff --git a/src/data/finsupp/order.lean b/src/data/finsupp/order.lean index 5d6ba134d6fc8..889ee49238192 100644 --- a/src/data/finsupp/order.lean +++ b/src/data/finsupp/order.lean @@ -8,6 +8,9 @@ import data.finsupp.defs /-! # Pointwise order on finitely supported functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file lifts order structures on `α` to `ι →₀ α`. ## Main declarations diff --git a/src/data/int/interval.lean b/src/data/int/interval.lean index eba687e822736..4856c9e400009 100644 --- a/src/data/int/interval.lean +++ b/src/data/int/interval.lean @@ -10,6 +10,9 @@ import data.finset.locally_finite /-! # Finite intervals of integers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves that `ℤ` is a `locally_finite_order` and calculates the cardinality of its intervals as finsets and fintypes. -/ diff --git a/src/data/list/cycle.lean b/src/data/list/cycle.lean index 8cfba07a9c6a6..5debeb8266bc0 100644 --- a/src/data/list/cycle.lean +++ b/src/data/list/cycle.lean @@ -10,6 +10,9 @@ import data.list.rotate /-! # Cycles of a list +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Lists have an equivalence relation of whether they are rotational permutations of one another. This relation is defined as `is_rotated`. diff --git a/src/data/nat/hyperoperation.lean b/src/data/nat/hyperoperation.lean index ca0f6f25a3b5e..739603bb2df95 100644 --- a/src/data/nat/hyperoperation.lean +++ b/src/data/nat/hyperoperation.lean @@ -9,6 +9,9 @@ import data.nat.parity /-! # Hyperoperation sequence +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the Hyperoperation sequence. `hyperoperation 0 m k = k + 1` `hyperoperation 1 m k = m + k` diff --git a/src/data/sym/sym2.lean b/src/data/sym/sym2.lean index f7287e12bd1cf..8e8f7af283968 100644 --- a/src/data/sym/sym2.lean +++ b/src/data/sym/sym2.lean @@ -11,6 +11,9 @@ import tactic.linarith /-! # The symmetric square +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the symmetric square, which is `α × α` modulo swapping. This is also known as the type of unordered pairs. diff --git a/src/deprecated/subgroup.lean b/src/deprecated/subgroup.lean index 2cd64ba456c5d..6e164389592e0 100644 --- a/src/deprecated/subgroup.lean +++ b/src/deprecated/subgroup.lean @@ -10,6 +10,9 @@ import deprecated.submonoid /-! # Unbundled subgroups (deprecated) +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file is deprecated, and is no longer imported by anything in mathlib other than other deprecated files, and test files. You should not need to import it. diff --git a/src/dynamics/periodic_pts.lean b/src/dynamics/periodic_pts.lean index bebf8b7ba3ac2..30e6211e65126 100644 --- a/src/dynamics/periodic_pts.lean +++ b/src/dynamics/periodic_pts.lean @@ -13,6 +13,9 @@ import group_theory.group_action.group /-! # Periodic points +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A point `x : α` is a periodic point of `f : α → α` of period `n` if `f^[n] x = x`. ## Main definitions diff --git a/src/group_theory/free_group.lean b/src/group_theory/free_group.lean index ed5d485a08040..d466c401c27f9 100644 --- a/src/group_theory/free_group.lean +++ b/src/group_theory/free_group.lean @@ -10,6 +10,9 @@ import group_theory.subgroup.basic /-! # Free groups +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines free groups over a type. Furthermore, it is shown that the free group construction is an instance of a monad. For the result that `free_group` is the left adjoint to the forgetful functor from groups to types, see `algebra/category/Group/adjunctions`. diff --git a/src/group_theory/group_action/conj_act.lean b/src/group_theory/group_action/conj_act.lean index 5a9ed7a2ffb8c..94d3aa0d6bdf4 100644 --- a/src/group_theory/group_action/conj_act.lean +++ b/src/group_theory/group_action/conj_act.lean @@ -9,6 +9,9 @@ import algebra.group_ring_action.basic /-! # Conjugation action of a group on itself +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the conjugation action of a group on itself. See also `mul_aut.conj` for the definition of conjugation as a homomorphism into the automorphism group. diff --git a/src/group_theory/group_action/fixing_subgroup.lean b/src/group_theory/group_action/fixing_subgroup.lean index ba7d078189e48..3f60d952629ee 100644 --- a/src/group_theory/group_action/fixing_subgroup.lean +++ b/src/group_theory/group_action/fixing_subgroup.lean @@ -11,6 +11,9 @@ import group_theory.group_action.basic # Fixing submonoid, fixing subgroup of an action +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In the presence of of an action of a monoid or a group, this file defines the fixing submonoid or the fixing subgroup, and relates it to the set of fixed points via a Galois connection. diff --git a/src/group_theory/group_action/sub_mul_action.lean b/src/group_theory/group_action/sub_mul_action.lean index 34f6cbcf8c2c5..dc5001117801e 100644 --- a/src/group_theory/group_action/sub_mul_action.lean +++ b/src/group_theory/group_action/sub_mul_action.lean @@ -11,6 +11,9 @@ import group_theory.group_action.basic # Sets invariant to a `mul_action` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define `sub_mul_action R M`; a subset of a `mul_action R M` which is closed with respect to scalar multiplication. diff --git a/src/group_theory/is_free_group.lean b/src/group_theory/is_free_group.lean index 528ee7e69965d..182863cda9e11 100644 --- a/src/group_theory/is_free_group.lean +++ b/src/group_theory/is_free_group.lean @@ -7,6 +7,9 @@ import group_theory.free_group /-! # Free groups structures on arbitrary types +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + 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. diff --git a/src/group_theory/perm/list.lean b/src/group_theory/perm/list.lean index 9b931cda52ffd..04aa0067202ee 100644 --- a/src/group_theory/perm/list.lean +++ b/src/group_theory/perm/list.lean @@ -10,6 +10,9 @@ import group_theory.perm.support /-! # Permutations from a list +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A list `l : list α` can be interpreted as a `equiv.perm α` where each element in the list is permuted to the next one, defined as `form_perm`. When we have that `nodup l`, we prove that `equiv.perm.support (form_perm l) = l.to_finset`, and that diff --git a/src/group_theory/subgroup/saturated.lean b/src/group_theory/subgroup/saturated.lean index e22b4921d7165..b3594a30926f0 100644 --- a/src/group_theory/subgroup/saturated.lean +++ b/src/group_theory/subgroup/saturated.lean @@ -9,6 +9,9 @@ import group_theory.subgroup.basic /-! # Saturated subgroups +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Tags subgroup, subgroups diff --git a/src/group_theory/subgroup/simple.lean b/src/group_theory/subgroup/simple.lean index f043081c9206e..b629a87b9ab54 100644 --- a/src/group_theory/subgroup/simple.lean +++ b/src/group_theory/subgroup/simple.lean @@ -9,6 +9,9 @@ import group_theory.subgroup.actions /-! # Simple groups +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines `is_simple_group G`, a class indicating that a group has exactly two normal subgroups. diff --git a/src/ring_theory/ring_invo.lean b/src/ring_theory/ring_invo.lean index 2d8a7bb246d2d..28fc739aec287 100644 --- a/src/ring_theory/ring_invo.lean +++ b/src/ring_theory/ring_invo.lean @@ -9,6 +9,9 @@ import algebra.ring.opposite /-! # Ring involutions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines a ring involution as a structure extending `R ≃+* Rᵐᵒᵖ`, with the additional fact `f.involution : (f (f x).unop).unop = x`. diff --git a/src/topology/bornology/hom.lean b/src/topology/bornology/hom.lean index f9b4a24cc1a80..021cac114f084 100644 --- a/src/topology/bornology/hom.lean +++ b/src/topology/bornology/hom.lean @@ -8,6 +8,9 @@ import topology.bornology.basic /-! # Locally bounded maps +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines locally bounded maps between bornologies. We use the `fun_like` design, so each type of morphisms has a companion typeclass which is meant to diff --git a/src/topology/locally_finite.lean b/src/topology/locally_finite.lean index 21bcd74e204d8..88b9bd2711022 100644 --- a/src/topology/locally_finite.lean +++ b/src/topology/locally_finite.lean @@ -9,6 +9,9 @@ import order.filter.small_sets /-! ### Locally finite families of sets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We say that a family of sets in a topological space is *locally finite* if at every point `x : X`, there is a neighborhood of `x` which meets only finitely many sets in the family. diff --git a/src/topology/nhds_set.lean b/src/topology/nhds_set.lean index 9f8adbbdb44aa..4a48a57b21136 100644 --- a/src/topology/nhds_set.lean +++ b/src/topology/nhds_set.lean @@ -7,6 +7,9 @@ import topology.basic /-! # Neighborhoods of a set +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define the filter `𝓝ˢ s` or `nhds_set s` consisting of all neighborhoods of a set `s`. From 48cb182b3de3950fc28bd6b603253a7222e9cb7f Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Mon, 30 Jan 2023 08:50:29 +0000 Subject: [PATCH 0381/1291] chore(ring_theory/fractional_ideal): remove instance name (#18305) --- src/ring_theory/fractional_ideal.lean | 3 +-- 1 file changed, 1 insertion(+), 2 deletions(-) diff --git a/src/ring_theory/fractional_ideal.lean b/src/ring_theory/fractional_ideal.lean index 7a28d8888942c..f24451c9a7637 100644 --- a/src/ring_theory/fractional_ideal.lean +++ b/src/ring_theory/fractional_ideal.lean @@ -912,8 +912,7 @@ lemma fractional_div_of_nonzero {I J : fractional_ideal R₁⁰ K} (h : J ≠ 0) I.is_fractional.div_of_nonzero J.is_fractional $ λ H, h $ coe_to_submodule_injective $ H.trans coe_zero.symm -noncomputable instance fractional_ideal_has_div : - has_div (fractional_ideal R₁⁰ K) := +noncomputable instance : has_div (fractional_ideal R₁⁰ K) := ⟨ λ I J, if h : J = 0 then 0 else ⟨I / J, fractional_div_of_nonzero h⟩ ⟩ variables {I J : fractional_ideal R₁⁰ K} [ J ≠ 0 ] From 68d1483e8a718ec63219f0e227ca3f0140361086 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Mon, 30 Jan 2023 15:11:08 +0000 Subject: [PATCH 0382/1291] chore(*): add mathlib4 synchronization comments (#18325) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.big_operators.intervals` * `algebra.free` * `algebra.geom_sum` * `algebra.hom.freiman` * `algebra.module.submodule.lattice` * `algebra.module.ulift` * `data.finset.slice` * `data.finsupp.fin` * `data.finsupp.indicator` * `data.finsupp.pointwise` * `data.num.bitwise` * `data.pnat.interval` * `data.set.pointwise.finite` * `data.set.pointwise.support` * `group_theory.semidirect_product` * `number_theory.divisors` * `order.disjointed` --- src/algebra/big_operators/intervals.lean | 3 +++ src/algebra/free.lean | 3 +++ src/algebra/geom_sum.lean | 3 +++ src/algebra/hom/freiman.lean | 3 +++ src/algebra/module/submodule/lattice.lean | 3 +++ src/algebra/module/ulift.lean | 3 +++ src/data/finset/slice.lean | 3 +++ src/data/finsupp/fin.lean | 3 +++ src/data/finsupp/indicator.lean | 3 +++ src/data/finsupp/pointwise.lean | 3 +++ src/data/num/bitwise.lean | 3 +++ src/data/pnat/interval.lean | 3 +++ src/data/set/pointwise/finite.lean | 5 ++++- src/data/set/pointwise/support.lean | 3 +++ src/group_theory/semidirect_product.lean | 3 +++ src/number_theory/divisors.lean | 3 +++ src/order/disjointed.lean | 3 +++ 17 files changed, 52 insertions(+), 1 deletion(-) diff --git a/src/algebra/big_operators/intervals.lean b/src/algebra/big_operators/intervals.lean index ac3699ea592de..32bf54f0133bf 100644 --- a/src/algebra/big_operators/intervals.lean +++ b/src/algebra/big_operators/intervals.lean @@ -12,6 +12,9 @@ import tactic.linarith /-! # Results about big operators over intervals +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We prove results about big operators over intervals (mostly the `ℕ`-valued `Ico m n`). -/ diff --git a/src/algebra/free.lean b/src/algebra/free.lean index dc7baa005232d..943813c801394 100644 --- a/src/algebra/free.lean +++ b/src/algebra/free.lean @@ -13,6 +13,9 @@ import data.list.basic /-! # Free constructions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions * `free_magma α`: free magma (structure with binary operation without any axioms) over alphabet `α`, diff --git a/src/algebra/geom_sum.lean b/src/algebra/geom_sum.lean index 2978020084607..69444d3774a38 100644 --- a/src/algebra/geom_sum.lean +++ b/src/algebra/geom_sum.lean @@ -13,6 +13,9 @@ import data.nat.parity /-! # Partial sums of geometric series +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file determines the values of the geometric series $\sum_{i=0}^{n-1} x^i$ and $\sum_{i=0}^{n-1} x^i y^{n-1-i}$ and variants thereof. We also provide some bounds on the "geometric" sum of `a/b^i` where `a b : ℕ`. diff --git a/src/algebra/hom/freiman.lean b/src/algebra/hom/freiman.lean index 3c32ac4ab79a4..08a9903902c69 100644 --- a/src/algebra/hom/freiman.lean +++ b/src/algebra/hom/freiman.lean @@ -9,6 +9,9 @@ import data.fun_like.basic /-! # Freiman homomorphisms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we define Freiman homomorphisms. A `n`-Freiman homomorphism on `A` is a function `f : α → β` such that `f (x₁) * ... * f (xₙ) = f (y₁) * ... * f (yₙ)` for all `x₁, ..., xₙ, y₁, ..., yₙ ∈ A` such that `x₁ * ... * xₙ = y₁ * ... * yₙ`. In particular, any diff --git a/src/algebra/module/submodule/lattice.lean b/src/algebra/module/submodule/lattice.lean index 8c1c341c0a73b..669c4f0ee22f9 100644 --- a/src/algebra/module/submodule/lattice.lean +++ b/src/algebra/module/submodule/lattice.lean @@ -9,6 +9,9 @@ import algebra.punit_instances /-! # The lattice structure on `submodule`s +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the lattice structure on submodules, `submodule.complete_lattice`, with `⊥` defined as `{0}` and `⊓` defined as intersection of the underlying carrier. If `p` and `q` are submodules of a module, `p ≤ q` means that `p ⊆ q`. diff --git a/src/algebra/module/ulift.lean b/src/algebra/module/ulift.lean index 1aeb5faba2016..cb9655b89cc93 100644 --- a/src/algebra/module/ulift.lean +++ b/src/algebra/module/ulift.lean @@ -9,6 +9,9 @@ import algebra.module.equiv /-! # `ulift` instances for module and multiplicative actions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines instances for module, mul_action and related structures on `ulift` types. (Recall `ulift α` is just a "copy" of a type `α` in a higher universe.) diff --git a/src/data/finset/slice.lean b/src/data/finset/slice.lean index 5014f7ed74cc5..b88806deb040e 100644 --- a/src/data/finset/slice.lean +++ b/src/data/finset/slice.lean @@ -10,6 +10,9 @@ import order.antichain /-! # `r`-sets and slice +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the `r`-th slice of a set family and provides a way to say that a set family is made of `r`-sets. diff --git a/src/data/finsupp/fin.lean b/src/data/finsupp/fin.lean index a8d468850acdd..eafd0136429cb 100644 --- a/src/data/finsupp/fin.lean +++ b/src/data/finsupp/fin.lean @@ -8,6 +8,9 @@ import data.finsupp.defs /-! # `cons` and `tail` for maps `fin n →₀ M` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We interpret maps `fin n →₀ M` as `n`-tuples of elements of `M`, We define the following operations: * `finsupp.tail` : the tail of a map `fin (n + 1) →₀ M`, i.e., its last `n` entries; diff --git a/src/data/finsupp/indicator.lean b/src/data/finsupp/indicator.lean index bf2a3f7ae1597..1cfe12ff0eb20 100644 --- a/src/data/finsupp/indicator.lean +++ b/src/data/finsupp/indicator.lean @@ -8,6 +8,9 @@ import data.finsupp.defs /-! # Building finitely supported functions off finsets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines `finsupp.indicator` to help create finsupps from finsets. ## Main declarations diff --git a/src/data/finsupp/pointwise.lean b/src/data/finsupp/pointwise.lean index 219f0bdabf84a..e2de0c10f122e 100644 --- a/src/data/finsupp/pointwise.lean +++ b/src/data/finsupp/pointwise.lean @@ -9,6 +9,9 @@ import algebra.ring.pi /-! # The pointwise product on `finsupp`. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + For the convolution product on `finsupp` when the domain has a binary operation, see the type synonyms `add_monoid_algebra` (which is in turn used to define `polynomial` and `mv_polynomial`) diff --git a/src/data/num/bitwise.lean b/src/data/num/bitwise.lean index d88123716468b..4ea3eb1ed0eb4 100644 --- a/src/data/num/bitwise.lean +++ b/src/data/num/bitwise.lean @@ -9,6 +9,9 @@ import data.bitvec.core /-! # Bitwise operations using binary representation of integers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Definitions * bitwise operations for `pos_num` and `num`, diff --git a/src/data/pnat/interval.lean b/src/data/pnat/interval.lean index fa6731fc6f2a6..6d2ecf63145a6 100644 --- a/src/data/pnat/interval.lean +++ b/src/data/pnat/interval.lean @@ -9,6 +9,9 @@ import data.pnat.defs /-! # Finite intervals of positive naturals +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves that `ℕ+` is a `locally_finite_order` and calculates the cardinality of its intervals as finsets and fintypes. -/ diff --git a/src/data/set/pointwise/finite.lean b/src/data/set/pointwise/finite.lean index 48814b4e7cc3f..cb402a79ba3ed 100644 --- a/src/data/set/pointwise/finite.lean +++ b/src/data/set/pointwise/finite.lean @@ -6,7 +6,10 @@ Authors: Johan Commelin, Floris van Doorn import data.set.finite import data.set.pointwise.smul -/-! # Finiteness lemmas for pointwise operations on sets -/ +/-! # Finiteness lemmas for pointwise operations on sets + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4.-/ open_locale pointwise diff --git a/src/data/set/pointwise/support.lean b/src/data/set/pointwise/support.lean index 2fe0996ee3a80..6d11a9c38352b 100644 --- a/src/data/set/pointwise/support.lean +++ b/src/data/set/pointwise/support.lean @@ -9,6 +9,9 @@ import algebra.support /-! # Support of a function composed with a scalar action +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We show that the support of `x ↦ f (c⁻¹ • x)` is equal to `c • support f`. -/ diff --git a/src/group_theory/semidirect_product.lean b/src/group_theory/semidirect_product.lean index 6ce04626068cc..86238ac809660 100644 --- a/src/group_theory/semidirect_product.lean +++ b/src/group_theory/semidirect_product.lean @@ -10,6 +10,9 @@ import group_theory.subgroup.basic /-! # Semidirect product +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines semidirect products of groups, and the canonical maps in and out of the semidirect product. The semidirect product of `N` and `G` given a hom `φ` from `G` to the automorphism group of `N` is the product of sets with the group diff --git a/src/number_theory/divisors.lean b/src/number_theory/divisors.lean index 3717371862f42..011394390da6d 100644 --- a/src/number_theory/divisors.lean +++ b/src/number_theory/divisors.lean @@ -10,6 +10,9 @@ import data.nat.factors /-! # Divisor finsets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines sets of divisors of a natural number. This is particularly useful as background for defining Dirichlet convolution. diff --git a/src/order/disjointed.lean b/src/order/disjointed.lean index 9b60b8f70bdea..ebb4a8605efd9 100644 --- a/src/order/disjointed.lean +++ b/src/order/disjointed.lean @@ -8,6 +8,9 @@ import order.partial_sups /-! # Consecutive differences of sets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the way to make a sequence of elements into a sequence of disjoint elements with the same partial sups. From 861a26926586cd46ff80264d121cdb6fa0e35cc1 Mon Sep 17 00:00:00 2001 From: Eric Rodriguez Date: Mon, 30 Jan 2023 18:01:35 +0000 Subject: [PATCH 0383/1291] chore(number_theory/cyclotomic/gal): tidy docs (#18326) Co-authored-by: Eric Rodriguez <37984851+ericrbg@users.noreply.github.com> --- src/number_theory/cyclotomic/gal.lean | 22 +++++++++++----------- 1 file changed, 11 insertions(+), 11 deletions(-) diff --git a/src/number_theory/cyclotomic/gal.lean b/src/number_theory/cyclotomic/gal.lean index 17ef9a05c2563..e9df88c21c521 100644 --- a/src/number_theory/cyclotomic/gal.lean +++ b/src/number_theory/cyclotomic/gal.lean @@ -17,10 +17,10 @@ it is always a subgroup, and if the `n`th cyclotomic polynomial is irreducible, * `is_primitive_root.aut_to_pow_injective`: `is_primitive_root.aut_to_pow` is injective in the case that it's considered over a cyclotomic field extension. -* `is_cyclotomic_extension.aut_equiv_pow`: If the `n`th cyclotomic polynomial is irreducible - in `K`, then `aut_to_pow` is a `mul_equiv` (for example, in `ℚ` and certain `𝔽ₚ`). -* `gal_X_pow_equiv_units_zmod`, `gal_cyclotomic_equiv_units_zmod`: Repackage `aut_equiv_pow` in - terms of `polynomial.gal`. +* `is_cyclotomic_extension.aut_equiv_pow`: If the `n`th cyclotomic polynomial is irreducible in `K`, + then `is_primitive_root.aut_to_pow` is a `mul_equiv` (for example, in `ℚ` and certain `𝔽ₚ`). +* `gal_X_pow_equiv_units_zmod`, `gal_cyclotomic_equiv_units_zmod`: Repackage + `is_cyclotomic_extension.aut_equiv_pow` in terms of `polynomial.gal`. * `is_cyclotomic_extension.aut.comm_group`: Cyclotomic extensions are abelian. ## References @@ -91,7 +91,7 @@ variables (h : irreducible (cyclotomic n K)) {K} (L) include h /-- The `mul_equiv` that takes an automorphism `f` to the element `k : (zmod n)ˣ` such that - `f μ = μ ^ k`. A stronger version of `is_primitive_root.aut_to_pow`. -/ + `f μ = μ ^ k` for any root of unity `μ`. A strengthening of `is_primitive_root.aut_to_pow`. -/ @[simps] noncomputable def aut_equiv_pow : (L ≃ₐ[K] L) ≃* (zmod n)ˣ := let hζ := zeta_spec n K L, hμ := λ t, hζ.pow_of_coprime _ (zmod.val_coe_unit_coprime t) in @@ -155,17 +155,17 @@ section gal variables [field L] (hμ : is_primitive_root μ n) [algebra K L] [is_cyclotomic_extension {n} K L] (h : irreducible (cyclotomic n K)) {K} -/-- `is_cyclotomic_extension.aut_equiv_pow` repackaged in terms of `gal`. Asserts that the -Galois group of `cyclotomic n K` is equivalent to `(zmod n)ˣ` if `cyclotomic n K` is irreducible in -the base field. -/ +/-- `is_cyclotomic_extension.aut_equiv_pow` repackaged in terms of `gal`. +Asserts that the Galois group of `cyclotomic n K` is equivalent to `(zmod n)ˣ` +if `cyclotomic n K` is irreducible in the base field. -/ noncomputable def gal_cyclotomic_equiv_units_zmod : (cyclotomic n K).gal ≃* (zmod n)ˣ := (alg_equiv.aut_congr (is_splitting_field.alg_equiv _ _)).symm.trans (is_cyclotomic_extension.aut_equiv_pow L h) -/-- `is_cyclotomic_extension.aut_equiv_pow` repackaged in terms of `gal`. Asserts that the -Galois group of `X ^ n - 1` is equivalent to `(zmod n)ˣ` if `cyclotomic n K` is irreducible in the -base field. -/ +/-- `is_cyclotomic_extension.aut_equiv_pow` repackaged in terms of `gal`. +Asserts that the Galois group of `X ^ n - 1` is equivalent to `(zmod n)ˣ` +if `cyclotomic n K` is irreducible in the base field. -/ noncomputable def gal_X_pow_equiv_units_zmod : (X ^ (n : ℕ) - 1).gal ≃* (zmod n)ˣ := (alg_equiv.aut_congr (is_splitting_field.alg_equiv _ _)).symm.trans From bcfa726826abd57587355b4b5b7e78ad6527b7e4 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Mon, 30 Jan 2023 21:46:11 +0000 Subject: [PATCH 0384/1291] refactor(topology/{order,*}): review API (#18312) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit ## Main changes * Switch from `@[class] structure topological_space` to `class`. * Introduce notation `is_open[t]`/`is_closed[t]`/`𝓤[u]` and use it instead of `t.is_open`/`@is_closed _ t`/`u.uniformity` * Don't introduce a temporary order on `topological_space`, use `galois_coinsertion` to the order-dual of `set (set α)` instead. * Drop `discrete_topology_bot`. It seems that this instance doesn't work well in Lean 4 (in fact, I failed to define it without using `@DiscreteTopology.mk`). ## Other changes ### Topological spaces * Add `is_open_mk`. * Rename `generate_from_mono` to `generate_from_anti`, move it to the `topological_space` namespace. * Add `embedding_inclusion`, `embedding_ulift_down`, and `ulift.discrete_topology`. * Move `discrete_topology.of_subset` from `topology.separation` to `topology.constructions`. * Move `embedding.discrete_topology` from `topology.separation` to `topology.maps`. * Use explicit arguments in an argument of `nhds_mk_of_nhds`. * Move some definitions and lemmas like `mk_of_closure`, `gi_generate_from` (renamed to `gci_generate_from`), `left_inverse_generate_from` to the `topological_space` namespace. These lemmas are very specific and use generic names. * Add `topological_space` and `discrete_topology` instances for `fin n`. * Drop `is_open_induced_iff'`, use non-primed version instead. * Prove `set_of_is_open_sup` by `rfl`. * Drop `nhds_bot`, use `nhds_discrete` instead. * Drop `induced_bot` and `discrete_topology_induced` ### Uniform spaces * Drop `infi_uniformity'` and `inf_uniformity'`. * Use `@uniformity α (uniform_space.comap _ _)` instead of `(h : ‹uniform_space α› = uniform_space.comap f ‹uniform_space β›)` in `uniformity_comap`. ### Locally constant functions and discrete quotients * Use quotient space topology for the coercion of a `discrete_quotient` to `Type*`, then prove `[discrete_topology]`. This way we avoid surprising diamonds in the future (especially if Lean 4 will unfold the coercion). * Merge `locally_constant.lift` and `locally_constant.locally_constant_lift` into 1 def, rename `factors` to `lift_comp_proj`. * Protect `locally_constant.is_locally_constant`. * Drop `locally_constant.iff_continuous_bot` ### Categories * Add an instance for `discrete_topology (discrete.obj X)`. * Rename `Fintype.discrete_topology` to `Fintype.bot_topology `, add lemma `Fintype.discrete_topology` sating that this is a `[discrete_topology]`. ### Topological rings * Fix&golf a proof that failed because of API changes. --- src/analysis/locally_convex/polar.lean | 3 +- src/analysis/normed_space/pi_Lp.lean | 2 +- src/geometry/manifold/charted_space.lean | 2 +- src/measure_theory/constructions/polish.lean | 12 +- src/topology/algebra/group/basic.lean | 3 +- src/topology/algebra/ring.lean | 21 +- src/topology/algebra/uniform_group.lean | 8 +- src/topology/bases.lean | 6 +- src/topology/basic.lean | 30 +- src/topology/category/CompHaus/basic.lean | 4 +- src/topology/category/Profinite/basic.lean | 14 +- src/topology/category/Top/basic.lean | 2 + src/topology/category/Top/limits.lean | 1 + src/topology/compact_open.lean | 4 +- src/topology/constructions.lean | 27 +- .../continuous_function/t0_sierpinski.lean | 4 +- src/topology/discrete_quotient.lean | 48 +-- src/topology/fiber_bundle/basic.lean | 2 +- src/topology/instances/discrete.lean | 2 +- src/topology/list.lean | 2 +- src/topology/locally_constant/algebra.lean | 2 +- src/topology/locally_constant/basic.lean | 8 +- src/topology/maps.lean | 7 + src/topology/metric_space/basic.lean | 9 +- src/topology/metric_space/emetric_space.lean | 4 +- src/topology/metric_space/polish.lean | 10 +- src/topology/order.lean | 345 ++++++++---------- src/topology/separation.lean | 32 +- src/topology/subset_properties.lean | 2 +- src/topology/uniform_space/basic.lean | 87 ++--- .../uniform_space/compact_convergence.lean | 4 +- src/topology/uniform_space/separation.lean | 6 +- .../uniform_convergence_topology.lean | 39 +- .../uniform_space/uniform_embedding.lean | 3 +- 34 files changed, 324 insertions(+), 431 deletions(-) diff --git a/src/analysis/locally_convex/polar.lean b/src/analysis/locally_convex/polar.lean index 1fdba000214d7..f02fa411c5415 100644 --- a/src/analysis/locally_convex/polar.lean +++ b/src/analysis/locally_convex/polar.lean @@ -37,6 +37,7 @@ polar variables {𝕜 E F : Type*} +open_locale topology namespace linear_map @@ -101,7 +102,7 @@ end /-- The polar set is closed in the weak topology induced by `B.flip`. -/ lemma polar_weak_closed (s : set E) : - @is_closed _ (weak_bilin.topological_space B.flip) (B.polar s) := + is_closed[weak_bilin.topological_space B.flip] (B.polar s) := begin rw polar_eq_Inter, refine is_closed_Inter (λ x, is_closed_Inter (λ _, _)), diff --git a/src/analysis/normed_space/pi_Lp.lean b/src/analysis/normed_space/pi_Lp.lean index 268e488b62f55..6c501a292c752 100644 --- a/src/analysis/normed_space/pi_Lp.lean +++ b/src/analysis/normed_space/pi_Lp.lean @@ -368,7 +368,7 @@ begin end lemma aux_uniformity_eq : - 𝓤 (pi_Lp p β) = @uniformity _ (Pi.uniform_space _) := + 𝓤 (pi_Lp p β) = 𝓤[Pi.uniform_space _] := begin have A : uniform_inducing (pi_Lp.equiv p β) := (antilipschitz_with_equiv_aux p β).uniform_inducing diff --git a/src/geometry/manifold/charted_space.lean b/src/geometry/manifold/charted_space.lean index 9584af912c1cb..4f25641a22c03 100644 --- a/src/geometry/manifold/charted_space.lean +++ b/src/geometry/manifold/charted_space.lean @@ -708,7 +708,7 @@ protected def to_topological_space : topological_space M := topological_space.generate_from $ ⋃ (e : local_equiv M H) (he : e ∈ c.atlas) (s : set H) (s_open : is_open s), {e ⁻¹' s ∩ e.source} -lemma open_source' (he : e ∈ c.atlas) : @is_open M c.to_topological_space e.source := +lemma open_source' (he : e ∈ c.atlas) : is_open[c.to_topological_space] e.source := begin apply topological_space.generate_open.basic, simp only [exists_prop, mem_Union, mem_singleton_iff], diff --git a/src/measure_theory/constructions/polish.lean b/src/measure_theory/constructions/polish.lean index 0e0655101819d..e04eadf44e41c 100644 --- a/src/measure_theory/constructions/polish.lean +++ b/src/measure_theory/constructions/polish.lean @@ -231,8 +231,8 @@ begin topology `t'`. It is analytic for this topology. As the identity from `t'` to `t` is continuous and the image of an analytic set is analytic, it follows that `s` is also analytic for `t`. -/ obtain ⟨t', t't, t'_polish, s_closed, s_open⟩ : - ∃ (t' : topological_space α), t' ≤ t ∧ @polish_space α t' ∧ @is_closed α t' s ∧ - @is_open α t' s := hs.is_clopenable, + ∃ t' : topological_space α, t' ≤ t ∧ @polish_space α t' ∧ is_closed[t'] s ∧ is_open[t'] s := + hs.is_clopenable, have A := @is_closed.analytic_set α t' t'_polish s s_closed, convert @analytic_set.image_of_continuous α t' α t s A id (continuous_id_of_le t't), simp only [id.def, image_id'], @@ -593,8 +593,8 @@ theorem _root_.measurable_set.image_of_continuous_on_inj_on measurable_set (f '' s) := begin obtain ⟨t', t't, t'_polish, s_closed, s_open⟩ : - ∃ (t' : topological_space γ), t' ≤ tγ ∧ @polish_space γ t' ∧ @is_closed γ t' s ∧ - @is_open γ t' s := hs.is_clopenable, + ∃ (t' : topological_space γ), t' ≤ tγ ∧ @polish_space γ t' ∧ is_closed[t'] s ∧ + is_open[t'] s := hs.is_clopenable, exact @is_closed.measurable_set_image_of_continuous_on_inj_on γ t' t'_polish β _ _ _ _ s s_closed f (f_cont.mono_dom t't) f_inj, end @@ -665,8 +665,8 @@ begin refine ⟨λ hs, _, λ hs, hs.is_clopenable⟩, -- consider a finer topology `t'` in which `s` is open and closed. obtain ⟨t', t't, t'_polish, s_closed, s_open⟩ : - ∃ (t' : topological_space γ), t' ≤ tγ ∧ @polish_space γ t' ∧ @is_closed γ t' s ∧ - @is_open γ t' s := hs, + ∃ (t' : topological_space γ), t' ≤ tγ ∧ @polish_space γ t' ∧ is_closed[t'] s ∧ + is_open[t'] s := hs, -- the identity is continuous from `t'` to `tγ`. have C : @continuous γ γ t' tγ id := continuous_id_of_le t't, -- therefore, it is also a measurable embedding, by the Lusin-Souslin theorem diff --git a/src/topology/algebra/group/basic.lean b/src/topology/algebra/group/basic.lean index e07a4c2f228eb..910414ccd7f62 100644 --- a/src/topology/algebra/group/basic.lean +++ b/src/topology/algebra/group/basic.lean @@ -1466,7 +1466,8 @@ instance : has_top (group_topology α) := @[to_additive] instance : has_bot (group_topology α) := ⟨{to_topological_space := ⊥, - continuous_mul := by continuity, + continuous_mul := by + { letI : topological_space α := ⊥, haveI := discrete_topology_bot α, continuity }, continuous_inv := continuous_bot}⟩ @[simp, to_additive] lemma to_topological_space_bot : diff --git a/src/topology/algebra/ring.lean b/src/topology/algebra/ring.lean index 899ca1d4bdd47..9ab1cc0bb9f1e 100644 --- a/src/topology/algebra/ring.lean +++ b/src/topology/algebra/ring.lean @@ -369,7 +369,7 @@ instance inhabited {α : Type u} [ring α] : inhabited (ring_topology α) := @[ext] lemma ext' {f g : ring_topology α} (h : f.is_open = g.is_open) : f = g := -by { ext, rw h } +by { ext : 2, exact h } /-- The ordering on ring topologies on the ring `α`. `t ≤ s` if every set open in `s` is also open in `t` (`t` is finer than `s`). -/ @@ -452,24 +452,7 @@ def to_add_group_topology (t : ring_topology α) : add_group_topology α := /-- The order embedding from ring topologies on `a` to additive group topologies on `a`. -/ def to_add_group_topology.order_embedding : order_embedding (ring_topology α) (add_group_topology α) := -{ to_fun := λ t, t.to_add_group_topology, - inj' := - begin - intros t₁ t₂ h_eq, - dsimp only at h_eq, - ext, - have h_t₁ : t₁.to_topological_space = t₁.to_add_group_topology.to_topological_space := rfl, - rw [h_t₁, h_eq], - refl, - end, - map_rel_iff' := - begin - intros t₁ t₂, - rw [embedding.coe_fn_mk], - have h_le : t₁ ≤ t₂ ↔ t₁.to_topological_space ≤ t₂.to_topological_space := by refl, - rw h_le, - refl, - end } +order_embedding.of_map_le_iff to_add_group_topology $ λ _ _, iff.rfl end ring_topology diff --git a/src/topology/algebra/uniform_group.lean b/src/topology/algebra/uniform_group.lean index d26baabfde3f8..7019aa3434da0 100644 --- a/src/topology/algebra/uniform_group.lean +++ b/src/topology/algebra/uniform_group.lean @@ -213,11 +213,8 @@ by { rw [← comap_swap_uniformity, uniformity_eq_comap_nhds_one, comap_comap, ( (hu : @uniform_group G u _) (hv : @uniform_group G v _) (h : @nhds _ u.to_topological_space 1 = @nhds _ v.to_topological_space 1) : u = v := -begin - refine uniform_space_eq _, - change @uniformity _ u = @uniformity _ v, - rw [@uniformity_eq_comap_nhds_one _ u _ hu, @uniformity_eq_comap_nhds_one _ v _ hv, h] -end +uniform_space_eq $ + by rw [@uniformity_eq_comap_nhds_one _ u _ hu, @uniformity_eq_comap_nhds_one _ v _ hv, h] @[to_additive] lemma uniform_group.ext_iff {G : Type*} [group G] {u v : uniform_space G} (hu : @uniform_group G u _) (hv : @uniform_group G v _) : @@ -579,7 +576,6 @@ end [group G] [uniform_group G] : topological_group.to_uniform_space G = u := begin ext : 1, - show @uniformity G (topological_group.to_uniform_space G) = 𝓤 G, rw [uniformity_eq_comap_nhds_one' G, uniformity_eq_comap_nhds_one G] end diff --git a/src/topology/bases.lean b/src/topology/bases.lean index 29294bf7158e4..cebad1c82cd8a 100644 --- a/src/topology/bases.lean +++ b/src/topology/bases.lean @@ -71,7 +71,7 @@ begin obtain ⟨t₃, h₃, hs⟩ := h.exists_subset_inter _ h₁ _ h₂ x ⟨hx₁, hx₂⟩, exact ⟨t₃, or.inr h₃, hs⟩ }, { rw h.eq_generate_from, - refine le_antisymm (le_generate_from $ λ t, _) (generate_from_mono $ subset_insert ∅ s), + refine le_antisymm (le_generate_from $ λ t, _) (generate_from_anti $ subset_insert ∅ s), rintro (rfl|ht), { convert is_open_empty }, { exact generate_open.basic t ht } }, end @@ -83,7 +83,7 @@ begin obtain ⟨t₃, h₃, hs⟩ := h.exists_subset_inter _ h₁ _ h₂ x hx, exact ⟨t₃, ⟨h₃, nonempty.ne_empty ⟨x, hs.1⟩⟩, hs⟩ }, { rw h.eq_generate_from, - refine le_antisymm (generate_from_mono $ diff_subset s _) (le_generate_from $ λ t ht, _), + refine le_antisymm (generate_from_anti $ diff_subset s _) (le_generate_from $ λ t ht, _), obtain rfl|he := eq_or_ne t ∅, { convert is_open_empty }, exact generate_open.basic t ⟨ht, he⟩ }, end @@ -93,7 +93,7 @@ subcollections of `s` form a topological basis. -/ lemma is_topological_basis_of_subbasis {s : set (set α)} (hs : t = generate_from s) : is_topological_basis ((λ f, ⋂₀ f) '' {f : set (set α) | f.finite ∧ f ⊆ s}) := begin - refine ⟨_, _, hs.trans (le_antisymm (le_generate_from _) $ generate_from_mono $ λ t ht, _)⟩, + refine ⟨_, _, hs.trans (le_antisymm (le_generate_from _) $ generate_from_anti $ λ t ht, _)⟩, { rintro _ ⟨t₁, ⟨hft₁, ht₁b⟩, rfl⟩ _ ⟨t₂, ⟨hft₂, ht₂b⟩, rfl⟩ x h, exact ⟨_, ⟨_, ⟨hft₁.union hft₂, union_subset ht₁b ht₂b⟩, sInter_union t₁ t₂⟩, h, subset.rfl⟩ }, { rw [sUnion_image, Union₂_eq_univ_iff], diff --git a/src/topology/basic.lean b/src/topology/basic.lean index 743bed8e70c69..c704a3fd2f161 100644 --- a/src/topology/basic.lean +++ b/src/topology/basic.lean @@ -62,14 +62,12 @@ universes u v w -/ /-- A topology on `α`. -/ -@[protect_proj] structure topological_space (α : Type u) := +@[protect_proj] class topological_space (α : Type u) := (is_open : set α → Prop) (is_open_univ : is_open univ) (is_open_inter : ∀s t, is_open s → is_open t → is_open (s ∩ t)) (is_open_sUnion : ∀s, (∀t∈s, is_open t) → is_open (⋃₀ s)) -attribute [class] topological_space - /-- A constructor for topologies by specifying the closed sets, and showing that they satisfy the appropriate conditions. -/ def topological_space.of_closed {α : Type u} (T : set (set α)) @@ -86,32 +84,36 @@ section topological_space variables {α : Type u} {β : Type v} {ι : Sort w} {a : α} {s s₁ s₂ t : set α} {p p₁ p₂ : α → Prop} +/-- `is_open s` means that `s` is open in the ambient topological space on `α` -/ +def is_open [topological_space α] (s : set α) : Prop := @topological_space.is_open _ ‹_› s + +localized "notation (name := is_open_of) `is_open[` t `]` := @is_open hole! t" in topology + +lemma is_open_mk {p h₁ h₂ h₃} {s : set α} : is_open[⟨p, h₁, h₂, h₃⟩] s ↔ p s := iff.rfl + @[ext] -lemma topological_space_eq : ∀ {f g : topological_space α}, f.is_open = g.is_open → f = g -| ⟨a, _, _, _⟩ ⟨b, _, _, _⟩ rfl := rfl +lemma topological_space_eq {f g : topological_space α} (h : is_open[f] = is_open[g]) : f = g := +by unfreezingI { cases f, cases g, congr, exact h } section variables [topological_space α] -/-- `is_open s` means that `s` is open in the ambient topological space on `α` -/ -def is_open (s : set α) : Prop := topological_space.is_open ‹_› s - @[simp] -lemma is_open_univ : is_open (univ : set α) := topological_space.is_open_univ _ +lemma is_open_univ : is_open (univ : set α) := topological_space.is_open_univ lemma is_open.inter (h₁ : is_open s₁) (h₂ : is_open s₂) : is_open (s₁ ∩ s₂) := -topological_space.is_open_inter _ s₁ s₂ h₁ h₂ +topological_space.is_open_inter s₁ s₂ h₁ h₂ lemma is_open_sUnion {s : set (set α)} (h : ∀t ∈ s, is_open t) : is_open (⋃₀ s) := -topological_space.is_open_sUnion _ s h +topological_space.is_open_sUnion s h end lemma topological_space_eq_iff {t t' : topological_space α} : - t = t' ↔ ∀ s, @is_open α t s ↔ @is_open α t' s := + t = t' ↔ ∀ s, is_open[t] s ↔ is_open[t'] s := ⟨λ h s, h ▸ iff.rfl, λ h, by { ext, exact h _ }⟩ -lemma is_open_fold {s : set α} {t : topological_space α} : t.is_open s = @is_open α t s := +lemma is_open_fold {s : set α} {t : topological_space α} : t.is_open s = is_open[t] s := rfl variables [topological_space α] @@ -165,6 +167,8 @@ is_open.inter class is_closed (s : set α) : Prop := (is_open_compl : is_open sᶜ) +localized "notation (name := is_closed_of) `is_closed[` t `]` := @is_closed hole! t" in topology + @[simp] lemma is_open_compl_iff {s : set α} : is_open sᶜ ↔ is_closed s := ⟨λ h, ⟨h⟩, λ h, h.is_open_compl⟩ diff --git a/src/topology/category/CompHaus/basic.lean b/src/topology/category/CompHaus/basic.lean index 44100f91149c3..2c7fa0f85ad20 100644 --- a/src/topology/category/CompHaus/basic.lean +++ b/src/topology/category/CompHaus/basic.lean @@ -235,8 +235,8 @@ lemma mono_iff_injective {X Y : CompHaus.{u}} (f : X ⟶ Y) : mono f ↔ functio begin split, { introsI hf x₁ x₂ h, - let g₁ : of punit ⟶ X := ⟨λ _, x₁, continuous_of_discrete_topology⟩, - let g₂ : of punit ⟶ X := ⟨λ _, x₂, continuous_of_discrete_topology⟩, + let g₁ : of punit ⟶ X := ⟨λ _, x₁, continuous_const⟩, + let g₂ : of punit ⟶ X := ⟨λ _, x₂, continuous_const⟩, have : g₁ ≫ f = g₂ ≫ f, by { ext, exact h }, rw cancel_mono at this, apply_fun (λ e, e punit.star) at this, diff --git a/src/topology/category/Profinite/basic.lean b/src/topology/category/Profinite/basic.lean index 91a665bc034c6..b07db2c8952d4 100644 --- a/src/topology/category/Profinite/basic.lean +++ b/src/topology/category/Profinite/basic.lean @@ -43,6 +43,7 @@ profinite universe u open category_theory +open_locale topology /-- The type of profinite topological spaces. -/ structure Profinite := @@ -134,11 +135,14 @@ lemma CompHaus.to_Profinite_obj' (X : CompHaus) : ↥(CompHaus.to_Profinite.obj X) = connected_components X := rfl /-- Finite types are given the discrete topology. -/ -def Fintype.discrete_topology (A : Fintype) : topological_space A := ⊥ +def Fintype.bot_topology (A : Fintype) : topological_space A := ⊥ section discrete_topology -local attribute [instance] Fintype.discrete_topology +local attribute [instance] Fintype.bot_topology + +local attribute [instance] +lemma Fintype.discrete_topology (A : Fintype) : discrete_topology A := ⟨rfl⟩ /-- The natural functor from `Fintype` to `Profinite`, endowing a finite type with the discrete topology. -/ @@ -254,17 +258,15 @@ lemma epi_iff_surjective {X Y : Profinite.{u}} (f : X ⟶ Y) : epi f ↔ functio begin split, { contrapose!, - rintros ⟨y, hy⟩ hf, + rintros ⟨y, hy⟩ hf, resetI, let C := set.range f, have hC : is_closed C := (is_compact_range f.continuous).is_closed, let U := Cᶜ, - have hU : is_open U := is_open_compl_iff.mpr hC, have hyU : y ∈ U, { refine set.mem_compl _, rintro ⟨y', hy'⟩, exact hy y' hy' }, - have hUy : U ∈ nhds y := hU.mem_nhds hyU, + have hUy : U ∈ 𝓝 y := hC.compl_mem_nhds hyU, obtain ⟨V, hV, hyV, hVU⟩ := is_topological_basis_clopen.mem_nhds_iff.mp hUy, classical, - letI : topological_space (ulift.{u} $ fin 2) := ⊥, let Z := of (ulift.{u} $ fin 2), let g : Y ⟶ Z := ⟨(locally_constant.of_clopen hV).map ulift.up, locally_constant.continuous _⟩, let h : Y ⟶ Z := ⟨λ _, ⟨1⟩, continuous_const⟩, diff --git a/src/topology/category/Top/basic.lean b/src/topology/category/Top/basic.lean index 561068bd765d9..5a2da0ff45f5c 100644 --- a/src/topology/category/Top/basic.lean +++ b/src/topology/category/Top/basic.lean @@ -55,6 +55,8 @@ def discrete : Type u ⥤ Top.{u} := { obj := λ X, ⟨X, ⊥⟩, map := λ X Y f, { to_fun := f, continuous_to_fun := continuous_bot } } +instance {X : Type u} : discrete_topology (discrete.obj X) := ⟨rfl⟩ + /-- The trivial topology on any type. -/ def trivial : Type u ⥤ Top.{u} := { obj := λ X, ⟨X, ⊤⟩, diff --git a/src/topology/category/Top/limits.lean b/src/topology/category/Top/limits.lean index 39b2769596412..d9d1b0c8925d4 100644 --- a/src/topology/category/Top/limits.lean +++ b/src/topology/category/Top/limits.lean @@ -1056,6 +1056,7 @@ lemma nonempty_sections_of_fintype_cofiltered_system.init F.sections.nonempty := begin let F' : J ⥤ Top := F ⋙ Top.discrete, + haveI : ∀ j, discrete_topology (F'.obj j) := λ _, ⟨rfl⟩, haveI : Π (j : J), fintype (F'.obj j) := hf, haveI : Π (j : J), nonempty (F'.obj j) := hne, obtain ⟨⟨u, hu⟩⟩ := Top.nonempty_limit_cone_of_compact_t2_cofiltered_system F', diff --git a/src/topology/compact_open.lean b/src/topology/compact_open.lean index 4fd8e8b260178..a1fb25e06a393 100644 --- a/src/topology/compact_open.lean +++ b/src/topology/compact_open.lean @@ -182,7 +182,7 @@ lemma compact_open_le_induced (s : set α) : ≤ topological_space.induced (continuous_map.restrict s) continuous_map.compact_open := begin simp only [induced_generate_from_eq, continuous_map.compact_open], - apply generate_from_mono, + apply topological_space.generate_from_anti, rintros b ⟨a, ⟨c, hc, u, hu, rfl⟩, rfl⟩, refine ⟨coe '' c, hc.image continuous_subtype_coe, u, hu, _⟩, ext f, @@ -202,7 +202,7 @@ begin { refine le_infi₂ _, exact λ s hs, compact_open_le_induced s }, simp only [← generate_from_Union, induced_generate_from_eq, continuous_map.compact_open], - apply generate_from_mono, + apply topological_space.generate_from_anti, rintros _ ⟨s, hs, u, hu, rfl⟩, rw mem_Union₂, refine ⟨s, hs, _, ⟨univ, is_compact_iff_is_compact_univ.mp hs, u, hu, rfl⟩, _⟩, diff --git a/src/topology/constructions.lean b/src/topology/constructions.lean index ca83960d42f7a..6e64a6c0ba3b8 100644 --- a/src/topology/constructions.lean +++ b/src/topology/constructions.lean @@ -466,8 +466,8 @@ lemma filter.has_basis.prod_nhds' {ιa ιb : Type*} {pa : ιa → Prop} {pb : ι by { cases ab, exact ha.prod_nhds hb } instance [discrete_topology α] [discrete_topology β] : discrete_topology (α × β) := -⟨eq_of_nhds_eq_nhds $ assume ⟨a, b⟩, - by rw [nhds_prod_eq, nhds_discrete α, nhds_discrete β, nhds_bot, filter.prod_pure_pure]⟩ +discrete_topology_iff_nhds.2 $ λ ⟨a, b⟩, + by rw [nhds_prod_eq, nhds_discrete α, nhds_discrete β, filter.prod_pure_pure] lemma prod_mem_nhds_iff {s : set α} {t : set β} {a : α} {b : β} : s ×ˢ t ∈ 𝓝 (a, b) ↔ s ∈ 𝓝 a ∧ t ∈ 𝓝 b := @@ -925,6 +925,16 @@ lemma embedding.cod_restrict {e : α → β} (he : embedding e) (s : set β) (hs embedding (cod_restrict e s hs) := embedding_of_embedding_compose (he.continuous.cod_restrict hs) continuous_subtype_coe he +lemma embedding_inclusion {s t : set α} (h : s ⊆ t) : embedding (set.inclusion h) := +embedding_subtype_coe.cod_restrict _ _ + +/-- Let `s, t ⊆ X` be two subsets of a topological space `X`. If `t ⊆ s` and the topology induced +by `X`on `s` is discrete, then also the topology induces on `t` is discrete. -/ +lemma discrete_topology.of_subset {X : Type*} [topological_space X] {s t : set X} + (ds : discrete_topology s) (ts : t ⊆ s) : + discrete_topology t := +(embedding_inclusion ts).discrete_topology + end subtype section quotient @@ -1089,7 +1099,7 @@ lemma pi_generate_from_eq {π : ι → Type*} {g : Πa, set (set (π a))} : let G := {t | ∃(s:Πa, set (π a)) (i : finset ι), (∀a∈i, s a ∈ g a) ∧ t = pi ↑i s} in begin rw [pi_eq_generate_from], - refine le_antisymm (generate_from_mono _) (le_generate_from _), + refine le_antisymm (generate_from_anti _) (le_generate_from _), exact assume s ⟨t, i, ht, eq⟩, ⟨t, i, assume a ha, generate_open.basic _ (ht a ha), eq⟩, { rintros s ⟨t, i, hi, rfl⟩, rw [pi_def], @@ -1106,8 +1116,8 @@ lemma pi_generate_from_eq_finite {π : ι → Type*} {g : Πa, set (set (π a))} begin casesI nonempty_fintype ι, rw [pi_generate_from_eq], - refine le_antisymm (generate_from_mono _) (le_generate_from _), - exact assume s ⟨t, ht, eq⟩, ⟨t, finset.univ, by simp [ht, eq]⟩, + refine le_antisymm (generate_from_anti _) (le_generate_from _), + { rintro s ⟨t, ht, rfl⟩, exact ⟨t, finset.univ, by simp [ht]⟩ }, { rintros s ⟨t, i, ht, rfl⟩, apply is_open_iff_forall_mem_open.2 _, assume f hf, @@ -1272,4 +1282,11 @@ continuous_induced_dom continuous (ulift.up : α → ulift.{v u} α) := continuous_induced_rng.2 continuous_id +lemma embedding_ulift_down [topological_space α] : + embedding (ulift.down : ulift.{v u} α → α) := +⟨⟨rfl⟩, ulift.down_injective⟩ + +instance [topological_space α] [discrete_topology α] : discrete_topology (ulift α) := +embedding_ulift_down.discrete_topology + end ulift diff --git a/src/topology/continuous_function/t0_sierpinski.lean b/src/topology/continuous_function/t0_sierpinski.lean index 6c3b386264580..a41e477f4054a 100644 --- a/src/topology/continuous_function/t0_sierpinski.lean +++ b/src/topology/continuous_function/t0_sierpinski.lean @@ -31,7 +31,7 @@ begin { intros u h, rw ← generate_from_Union_is_open, apply is_open_generate_from_of_mem, - simp only [set.mem_Union, set.mem_set_of_eq, is_open_induced_iff'], + simp only [set.mem_Union, set.mem_set_of_eq, is_open_induced_iff], exact ⟨⟨u, h⟩, {true}, is_open_singleton_true, by simp [set.preimage]⟩ }, end @@ -41,7 +41,7 @@ variables (X : Type*) [topological_space X] The continuous map from `X` to the product of copies of the Sierpinski space, (one copy for each open subset `u` of `X`). The `u` coordinate of `product_of_mem_opens x` is given by `x ∈ u`. -/ -def product_of_mem_opens : continuous_map X (opens X → Prop) := +def product_of_mem_opens : C(X, opens X → Prop) := { to_fun := λ x u, x ∈ u, continuous_to_fun := continuous_pi_iff.2 (λ u, continuous_Prop.2 u.property) } diff --git a/src/topology/discrete_quotient.lean b/src/topology/discrete_quotient.lean index 16cc4543bb6fd..cdd6ce05a6e50 100644 --- a/src/topology/discrete_quotient.lean +++ b/src/topology/discrete_quotient.lean @@ -95,13 +95,11 @@ def setoid : setoid X := ⟨S.rel, S.equiv⟩ instance : has_coe_to_sort (discrete_quotient X) Type* := ⟨λ S, quotient S.setoid⟩ -instance : topological_space S := ⊥ +instance : topological_space S := quotient.topological_space /-- The projection from `X` to the given discrete quotient. -/ def proj : X → S := quotient.mk' -lemma proj_surjective : function.surjective S.proj := quotient.surjective_quotient_mk' - lemma fiber_eq (x : X) : S.proj ⁻¹' {S.proj x} = set_of (S.rel x) := begin ext1 y, @@ -110,24 +108,22 @@ begin exact ⟨λ h, S.symm _ _ h, λ h, S.symm _ _ h⟩, end -lemma proj_is_locally_constant : is_locally_constant S.proj := -begin - rw (is_locally_constant.tfae S.proj).out 0 3, - intros x, - rcases S.proj_surjective x with ⟨x,rfl⟩, - simp [fiber_eq, (S.clopen x).1], -end +lemma proj_surjective : function.surjective S.proj := quotient.surjective_quotient_mk' +lemma proj_quotient_map : quotient_map S.proj := quotient_map_quot_mk +lemma proj_continuous : continuous S.proj := S.proj_quotient_map.continuous -lemma proj_continuous : continuous S.proj := -is_locally_constant.continuous $ proj_is_locally_constant _ +instance : discrete_topology S := +singletons_open_iff_discrete.1 $ S.proj_surjective.forall.2 $ λ x, + by { rw [← S.proj_quotient_map.is_open_preimage, fiber_eq], exact (S.clopen _).1 } -lemma fiber_closed (A : set S) : is_closed (S.proj ⁻¹' A) := -is_closed.preimage S.proj_continuous ⟨trivial⟩ +lemma proj_is_locally_constant : is_locally_constant S.proj := +(is_locally_constant.iff_continuous S.proj).2 S.proj_continuous -lemma fiber_open (A : set S) : is_open (S.proj ⁻¹' A) := -is_open.preimage S.proj_continuous trivial +lemma fiber_clopen (A : set S) : is_clopen (S.proj ⁻¹' A) := +(is_clopen_discrete A).preimage S.proj_continuous -lemma fiber_clopen (A : set S) : is_clopen (S.proj ⁻¹' A) := ⟨fiber_open _ _, fiber_closed _ _⟩ +lemma fiber_open (A : set S) : is_open (S.proj ⁻¹' A) := (S.fiber_clopen A).1 +lemma fiber_closed (A : set S) : is_closed (S.proj ⁻¹' A) := (S.fiber_clopen A).2 instance : partial_order (discrete_quotient X) := { le := λ A B, ∀ x y : X, A.rel x y → B.rel x y, @@ -355,19 +351,11 @@ def discrete_quotient : discrete_quotient X := equiv := ⟨by tauto, by tauto, λ a b c h1 h2, by rw [h2, h1]⟩, clopen := λ x, f.is_locally_constant.is_clopen_fiber _ } -/-- The function from the discrete quotient associated to a locally constant function. -/ -def lift : f.discrete_quotient → α := λ a, quotient.lift_on' a f (λ a b h, h.symm) - -lemma lift_is_locally_constant : _root_.is_locally_constant f.lift := λ A, trivial - -/-- A locally constant version of `locally_constant.lift`. -/ -def locally_constant_lift : locally_constant f.discrete_quotient α := -⟨f.lift, f.lift_is_locally_constant⟩ +/-- The (locally constant) function from the discrete quotient associated to a locally constant +function. -/ +def lift : locally_constant f.discrete_quotient α := +⟨λ a, quotient.lift_on' a f (λ a b h, h.symm), λ A, is_open_discrete _⟩ -@[simp] -lemma lift_eq_coe : f.lift = f.locally_constant_lift := rfl - -@[simp] -lemma factors : f.locally_constant_lift ∘ f.discrete_quotient.proj = f := by { ext, refl } +@[simp] lemma lift_comp_proj : f.lift ∘ f.discrete_quotient.proj = f := by { ext, refl } end locally_constant diff --git a/src/topology/fiber_bundle/basic.lean b/src/topology/fiber_bundle/basic.lean index 62cfece2a4958..43445113cbc46 100644 --- a/src/topology/fiber_bundle/basic.lean +++ b/src/topology/fiber_bundle/basic.lean @@ -693,7 +693,7 @@ begin exact le_supr₂ e he, end -lemma is_open_source (e : pretrivialization F (π E)) : @is_open _ a.total_space_topology e.source := +lemma is_open_source (e : pretrivialization F (π E)) : is_open[a.total_space_topology] e.source := begin letI := a.total_space_topology, refine is_open_supr_iff.mpr (λ e', _), diff --git a/src/topology/instances/discrete.lean b/src/topology/instances/discrete.lean index 1dabbe8d159fb..3bb76457acdbc 100644 --- a/src/topology/instances/discrete.lean +++ b/src/topology/instances/discrete.lean @@ -126,6 +126,6 @@ discrete_topology_iff_order_topology_of_pred_succ.mp h @[priority 100] instance discrete_topology.metrizable_space [discrete_topology α] : metrizable_space α := begin - rw discrete_topology.eq_bot α, + unfreezingI { obtain rfl := discrete_topology.eq_bot α }, exact @uniform_space.metrizable_space α ⊥ (is_countably_generated_principal _) _, end diff --git a/src/topology/list.lean b/src/topology/list.lean index a7a761a61516b..0626f4cdf25c7 100644 --- a/src/topology/list.lean +++ b/src/topology/list.lean @@ -34,7 +34,7 @@ begin { existsi [], simpa only [list.forall₂_nil_left_iff, exists_eq_left] }, case list.forall₂.cons : a s as ss ht h ih t hts { rcases mem_nhds_iff.1 ht with ⟨u, hut, hu⟩, - rcases ih (subset.refl _) with ⟨v, hv, hvss⟩, + rcases ih _ subset.rfl with ⟨v, hv, hvss⟩, exact ⟨u::v, list.forall₂.cons hu hv, subset.trans (set.seq_mono (set.image_subset _ hut) hvss) hts⟩ } }, rcases this with ⟨v, hv, hvs⟩, diff --git a/src/topology/locally_constant/algebra.lean b/src/topology/locally_constant/algebra.lean index 3f1eda30bc7b0..740b15593dac6 100644 --- a/src/topology/locally_constant/algebra.lean +++ b/src/topology/locally_constant/algebra.lean @@ -192,7 +192,7 @@ variables {R : Type*} instance [has_smul R Y] : has_smul R (locally_constant X Y) := { smul := λ r f, { to_fun := r • f, - is_locally_constant := ((is_locally_constant f).comp ((•) r) : _), } } + is_locally_constant := (f.is_locally_constant.comp ((•) r) : _), } } @[simp] lemma coe_smul [has_smul R Y] (r : R) (f : locally_constant X Y) : ⇑(r • f) = r • f := rfl diff --git a/src/topology/locally_constant/basic.lean b/src/topology/locally_constant/basic.lean index 84cbb39b4d66c..436f30749555f 100644 --- a/src/topology/locally_constant/basic.lean +++ b/src/topology/locally_constant/basic.lean @@ -97,10 +97,6 @@ lemma iff_continuous {_ : topological_space Y} [discrete_topology Y] (f : X → is_locally_constant f ↔ continuous f := ⟨is_locally_constant.continuous, λ h s, h.is_open_preimage s (is_open_discrete _)⟩ -lemma iff_continuous_bot (f : X → Y) : - is_locally_constant f ↔ @continuous X Y _ ⊥ f := -iff_continuous f - lemma of_constant (f : X → Y) (h : ∀ x y, f x = f y) : is_locally_constant f := (iff_eventually_eq f).2 $ λ x, eventually_of_forall $ λ x', h _ _ @@ -165,8 +161,7 @@ lemma iff_is_const [preconnected_space X] {f : X → Y} : lemma range_finite [compact_space X] {f : X → Y} (hf : is_locally_constant f) : (set.range f).finite := begin - letI : topological_space Y := ⊥, - haveI : discrete_topology Y := ⟨rfl⟩, + letI : topological_space Y := ⊥, haveI := discrete_topology_bot Y, rw @iff_continuous X Y ‹_› ‹_› at hf, exact (is_compact_range hf).finite_of_discrete end @@ -219,6 +214,7 @@ of_constant_on_connected_components (λ x, h (connected_component x) end is_locally_constant /-- A (bundled) locally constant function from a topological space `X` to a type `Y`. -/ +@[protect_proj] structure locally_constant (X Y : Type*) [topological_space X] := (to_fun : X → Y) (is_locally_constant : is_locally_constant to_fun) diff --git a/src/topology/maps.lean b/src/topology/maps.lean index 38cc9151f2ec3..3b3cc60f64cfb 100644 --- a/src/topology/maps.lean +++ b/src/topology/maps.lean @@ -210,6 +210,13 @@ lemma embedding.closure_eq_preimage_closure_image {e : α → β} (he : embeddin closure s = e ⁻¹' closure (e '' s) := he.1.closure_eq_preimage_closure_image s +/-- The topology induced under an inclusion `f : X → Y` from the discrete topological space `Y` +is the discrete topology on `X`. -/ +lemma embedding.discrete_topology {X Y : Type*} [topological_space X] [tY : topological_space Y] + [discrete_topology Y] {f : X → Y} (hf : embedding f) : discrete_topology X := +discrete_topology_iff_nhds.2 $ λ x, by rw [hf.nhds_eq_comap, nhds_discrete, comap_pure, + ← image_singleton, hf.inj.preimage_image, principal_singleton] + end embedding /-- A function between topological spaces is a quotient map if it is surjective, diff --git a/src/topology/metric_space/basic.lean b/src/topology/metric_space/basic.lean index 68d8dbc87ee33..c1364354d899c 100644 --- a/src/topology/metric_space/basic.lean +++ b/src/topology/metric_space/basic.lean @@ -1127,7 +1127,7 @@ by rw [emetric.inseparable_iff, edist_nndist, dist_nndist, ennreal.coe_eq_zero, See Note [forgetful inheritance]. -/ def pseudo_metric_space.replace_uniformity {α} [U : uniform_space α] (m : pseudo_metric_space α) - (H : @uniformity _ U = @uniformity _ pseudo_emetric_space.to_uniform_space) : + (H : 𝓤[U] = 𝓤[pseudo_emetric_space.to_uniform_space]) : pseudo_metric_space α := { dist := @dist _ m.to_has_dist, dist_self := dist_self, @@ -1139,8 +1139,7 @@ def pseudo_metric_space.replace_uniformity {α} [U : uniform_space α] (m : pseu uniformity_dist := H.trans pseudo_metric_space.uniformity_dist } lemma pseudo_metric_space.replace_uniformity_eq {α} [U : uniform_space α] - (m : pseudo_metric_space α) - (H : @uniformity _ U = @uniformity _ pseudo_emetric_space.to_uniform_space) : + (m : pseudo_metric_space α) (H : 𝓤[U] = 𝓤[pseudo_emetric_space.to_uniform_space]) : m.replace_uniformity H = m := by { ext, refl } @@ -2769,13 +2768,13 @@ end metric See Note [forgetful inheritance]. -/ def metric_space.replace_uniformity {γ} [U : uniform_space γ] (m : metric_space γ) - (H : @uniformity _ U = @uniformity _ pseudo_emetric_space.to_uniform_space) : + (H : 𝓤[U] = 𝓤[pseudo_emetric_space.to_uniform_space]) : metric_space γ := { eq_of_dist_eq_zero := @eq_of_dist_eq_zero _ _, ..pseudo_metric_space.replace_uniformity m.to_pseudo_metric_space H, } lemma metric_space.replace_uniformity_eq {γ} [U : uniform_space γ] (m : metric_space γ) - (H : @uniformity _ U = @uniformity _ pseudo_emetric_space.to_uniform_space) : + (H : 𝓤[U] = 𝓤[pseudo_emetric_space.to_uniform_space]) : m.replace_uniformity H = m := by { ext, refl } diff --git a/src/topology/metric_space/emetric_space.lean b/src/topology/metric_space/emetric_space.lean index cbbd0a29d972f..275a381e11557 100644 --- a/src/topology/metric_space/emetric_space.lean +++ b/src/topology/metric_space/emetric_space.lean @@ -377,7 +377,7 @@ specified uniformity. See Note [forgetful inheritance] explaining why having def the right uniformity is often important. -/ def pseudo_emetric_space.replace_uniformity {α} [U : uniform_space α] (m : pseudo_emetric_space α) - (H : @uniformity _ U = @uniformity _ pseudo_emetric_space.to_uniform_space) : + (H : 𝓤[U] = 𝓤[pseudo_emetric_space.to_uniform_space]) : pseudo_emetric_space α := { edist := @edist _ m.to_has_edist, edist_self := edist_self, @@ -959,7 +959,7 @@ specified uniformity. See Note [forgetful inheritance] explaining why having def the right uniformity is often important. -/ def emetric_space.replace_uniformity {γ} [U : uniform_space γ] (m : emetric_space γ) - (H : @uniformity _ U = @uniformity _ pseudo_emetric_space.to_uniform_space) : + (H : 𝓤[U] = 𝓤[pseudo_emetric_space.to_uniform_space]) : emetric_space γ := { edist := @edist _ m.to_has_edist, edist_self := edist_self, diff --git a/src/topology/metric_space/polish.lean b/src/topology/metric_space/polish.lean index 253f5d1a46f89..cfcf50428ebda 100644 --- a/src/topology/metric_space/polish.lean +++ b/src/topology/metric_space/polish.lean @@ -413,7 +413,7 @@ end complete_copy this set is open and closed. It turns out that this notion is equivalent to being Borel-measurable, but this is nontrivial (see `is_clopenable_iff_measurable_set`). -/ def is_clopenable [t : topological_space α] (s : set α) : Prop := -∃ (t' : topological_space α), t' ≤ t ∧ @polish_space α t' ∧ @is_closed α t' s ∧ @is_open α t' s +∃ (t' : topological_space α), t' ≤ t ∧ @polish_space α t' ∧ is_closed[t'] s ∧ is_open[t'] s /-- Given a closed set `s` in a Polish space, one can construct a finer Polish topology for which `s` is both open and closed. -/ @@ -450,13 +450,13 @@ begin have : sum.inr ⁻¹' (⇑(f.symm) ⁻¹' u) = (coe : t → α) ⁻¹' u, by { ext x, simp only [equiv.symm_symm, mem_preimage, equiv.set.sum_compl_apply_inr] }, rwa this } }, - { have : @is_closed α t' (g ⁻¹' (range (sum.inl : s → s ⊕ t))), + { have : is_closed[t'] (g ⁻¹' (range (sum.inl : s → s ⊕ t))), { apply is_closed.preimage, { exact @homeomorph.continuous _ _ t' _ g }, { exact is_closed_range_inl } }, convert this, exact A.symm }, - { have : @is_open α t' (g ⁻¹' (range (sum.inl : s → s ⊕ t))), + { have : is_open[t'] (g ⁻¹' (range (sum.inl : s → s ⊕ t))), { apply is_open.preimage, { exact @homeomorph.continuous _ _ t' _ g }, { exact is_open_range_inl } }, @@ -483,14 +483,14 @@ begin obtain ⟨t', t'm, -, t'_polish⟩ : ∃ (t' : topological_space α), (∀ (n : ℕ), t' ≤ m n) ∧ (t' ≤ t) ∧ @polish_space α t' := exists_polish_space_forall_le m mt m_polish, - have A : @is_open α t' (⋃ n, s n), + have A : is_open[t'] (⋃ n, s n), { apply is_open_Union, assume n, apply t'm n, exact m_open n }, obtain ⟨t'', t''_le, t''_polish, h1, h2⟩ : ∃ (t'' : topological_space α), t'' ≤ t' ∧ @polish_space α t'' - ∧ @is_closed α t'' (⋃ n, s n) ∧ @is_open α t'' (⋃ n, s n) := + ∧ is_closed[t''] (⋃ n, s n) ∧ is_open[t''] (⋃ n, s n) := @is_open.is_clopenable α t' t'_polish _ A, exact ⟨t'', t''_le.trans ((t'm 0).trans (mt 0)), t''_polish, h1, h2⟩, end diff --git a/src/topology/order.lean b/src/topology/order.lean index 03acaaacd2b62..04914a90301e1 100644 --- a/src/topology/order.lean +++ b/src/topology/order.lean @@ -42,7 +42,7 @@ finer, coarser, induced topology, coinduced topology -/ open function set filter -open_locale classical topology filter +open_locale topology filter universes u v w @@ -64,28 +64,23 @@ def generate_from (g : set (set α)) : topological_space α := is_open_sUnion := generate_open.sUnion } lemma is_open_generate_from_of_mem {g : set (set α)} {s : set α} (hs : s ∈ g) : - @is_open _ (generate_from g) s := + is_open[generate_from g] s := generate_open.basic s hs lemma nhds_generate_from {g : set (set α)} {a : α} : @nhds α (generate_from g) a = (⨅s∈{s | a ∈ s ∧ s ∈ g}, 𝓟 s) := -by rw nhds_def; exact le_antisymm - (binfi_mono $ λ s ⟨as, sg⟩, ⟨as, generate_open.basic _ sg⟩) - (le_infi $ assume s, le_infi $ assume ⟨as, hs⟩, - begin - revert as, clear_, induction hs, - case generate_open.basic : s hs - { exact assume as, infi_le_of_le s $ infi_le _ ⟨as, hs⟩ }, - case generate_open.univ - { rw [principal_univ], - exact assume _, le_top }, - case generate_open.inter : s t hs' ht' hs ht - { exact assume ⟨has, hat⟩, calc _ ≤ 𝓟 s ⊓ 𝓟 t : le_inf (hs has) (ht hat) - ... = _ : inf_principal }, - case generate_open.sUnion : k hk' hk - { exact λ ⟨t, htk, hat⟩, calc _ ≤ 𝓟 t : hk t htk hat - ... ≤ _ : le_principal_iff.2 $ subset_sUnion_of_mem htk } - end) +begin + rw nhds_def, + refine le_antisymm (binfi_mono $ λ s ⟨as, sg⟩, ⟨as, generate_open.basic _ sg⟩) _, + refine le_infi₂ (λ s hs, _), cases hs with ha hs, + induction hs, + case basic : s hs { exact infi₂_le _ ⟨ha, hs⟩ }, + case univ : { exact le_top.trans_eq principal_univ.symm }, + case inter : s t hs' ht' hs ht { exact (le_inf (hs ha.1) (ht ha.2)).trans_eq inf_principal }, + case sUnion : S hS' hS + { rcases ha with ⟨t, htS, hat⟩, + exact (hS t htS hat).trans (principal_mono.2 $ subset_sUnion_of_mem htS) } +end lemma tendsto_nhds_generate_from {β : Type*} {m : α → β} {f : filter α} {g : set (set β)} {b : β} (h : ∀s∈g, b ∈ s → m ⁻¹' s ∈ f) : tendsto m f (@nhds β (generate_from g) b) := @@ -101,7 +96,7 @@ protected def mk_of_nhds (n : α → filter α) : topological_space α := mem_of_superset (hs x hx _ hxa) (set.subset_sUnion_of_mem hx) } lemma nhds_mk_of_nhds (n : α → filter α) (a : α) - (h₀ : pure ≤ n) (h₁ : ∀{a s}, s ∈ n a → ∃ t ∈ n a, t ⊆ s ∧ ∀a' ∈ t, s ∈ n a') : + (h₀ : pure ≤ n) (h₁ : ∀ a s, s ∈ n a → ∃ t ∈ n a, t ⊆ s ∧ ∀a' ∈ t, s ∈ n a') : @nhds α (topological_space.mk_of_nhds n) a = n a := begin letI := topological_space.mk_of_nhds n, @@ -109,7 +104,7 @@ begin { have h₀ : {b | s ∈ n b} ⊆ s := assume b hb, mem_pure.1 $ h₀ b hb, have h₁ : {b | s ∈ n b} ∈ 𝓝 a, { refine is_open.mem_nhds (assume b (hb : s ∈ n b), _) hs, - rcases h₁ hb with ⟨t, ht, hts, h⟩, + rcases h₁ _ _ hb with ⟨t, ht, hts, h⟩, exact mem_of_superset ht h }, exact mem_of_superset h₁ h₀ }, { rcases (@mem_nhds_iff α (topological_space.mk_of_nhds n) _ _).1 hs with ⟨t, hts, ht, hat⟩, @@ -145,120 +140,100 @@ begin exact ⟨n₂, hn₄, hn₅.trans hm₂⟩, }, end -end topological_space - section lattice -variables {α : Type u} {β : Type v} +/-- The ordering on topologies on the type `α`. `t ≤ s` if every set open in `s` is also open in `t` +(`t` is finer than `s`). -/ +instance : partial_order (topological_space α) := +{ le := λ s t, ∀ U, is_open[t] U → is_open[s] U, + .. partial_order.lift (λ s, order_dual.to_dual (is_open[s])) (λ _ _, topological_space_eq) } -/-- The inclusion ordering on topologies on α. We use it to get a complete - lattice instance via the Galois insertion method, but the partial order - that we will eventually impose on `topological_space α` is the reverse one. -/ -def tmp_order : partial_order (topological_space α) := -{ le := λt s, t.is_open ≤ s.is_open, - le_antisymm := assume t s h₁ h₂, topological_space_eq $ le_antisymm h₁ h₂, - le_refl := assume t, le_refl t.is_open, - le_trans := assume a b c h₁ h₂, @le_trans _ _ a.is_open b.is_open c.is_open h₁ h₂ } - -local attribute [instance] tmp_order - -/- We'll later restate this lemma in terms of the correct order on `topological_space α`. -/ -private lemma generate_from_le_iff_subset_is_open {g : set (set α)} {t : topological_space α} : - topological_space.generate_from g ≤ t ↔ g ⊆ {s | t.is_open s} := -iff.intro - (assume ht s hs, ht _ $ topological_space.generate_open.basic s hs) - (assume hg s hs, hs.rec_on (assume v hv, hg hv) - t.is_open_univ (assume u v _ _, t.is_open_inter u v) (assume k _, t.is_open_sUnion k)) - -/-- If `s` equals the collection of open sets in the topology it generates, - then `s` defines a topology. -/ -protected def mk_of_closure (s : set (set α)) - (hs : {u | (topological_space.generate_from s).is_open u} = s) : topological_space α := -{ is_open := λu, u ∈ s, +protected lemma le_def {α} {t s : topological_space α} : t ≤ s ↔ is_open[s] ≤ is_open[t] := +iff.rfl + +lemma le_generate_from_iff_subset_is_open {g : set (set α)} {t : topological_space α} : + t ≤ topological_space.generate_from g ↔ g ⊆ {s | is_open[t] s} := +⟨λ ht s hs, ht _ $ generate_open.basic s hs, λ hg s hs, hs.rec_on (assume v hv, hg hv) + t.is_open_univ (assume u v _ _, t.is_open_inter u v) (assume k _, t.is_open_sUnion k)⟩ + +/-- If `s` equals the collection of open sets in the topology it generates, then `s` defines a +topology. -/ +protected def mk_of_closure (s : set (set α)) (hs : {u | generate_open s u} = s) : + topological_space α := +{ is_open := λ u, u ∈ s, is_open_univ := hs ▸ topological_space.generate_open.univ, is_open_inter := hs ▸ topological_space.generate_open.inter, is_open_sUnion := hs ▸ topological_space.generate_open.sUnion } -lemma mk_of_closure_sets {s : set (set α)} - {hs : {u | (topological_space.generate_from s).is_open u} = s} : - mk_of_closure s hs = topological_space.generate_from s := +lemma mk_of_closure_sets {s : set (set α)} {hs : {u | generate_open s u} = s} : + topological_space.mk_of_closure s hs = topological_space.generate_from s := topological_space_eq hs.symm -/-- The Galois insertion between `set (set α)` and `topological_space α` whose lower part - sends a collection of subsets of α to the topology they generate, and whose upper part - sends a topology to its collection of open subsets. -/ -def gi_generate_from (α : Type*) : - galois_insertion topological_space.generate_from (λt:topological_space α, {s | t.is_open s}) := -{ gc := assume g t, generate_from_le_iff_subset_is_open, - le_l_u := assume ts s hs, topological_space.generate_open.basic s hs, - choice := λg hg, mk_of_closure g - (subset.antisymm hg $ generate_from_le_iff_subset_is_open.1 $ le_rfl), +lemma gc_generate_from (α) : + galois_connection (λ t : topological_space α, order_dual.to_dual {s | is_open[t] s}) + (generate_from ∘ order_dual.of_dual) := +λ _ _, le_generate_from_iff_subset_is_open.symm + +/-- The Galois coinsertion between `topological_space α` and `(set (set α))ᵒᵈ` whose lower part + sends a topology to its collection of open subsets, and whose upper part sends a collection of + subsets of α to the topology they generate. -/ +def gci_generate_from (α : Type*) : + galois_coinsertion (λ t : topological_space α, order_dual.to_dual {s | is_open[t] s}) + (generate_from ∘ order_dual.of_dual) := +{ gc := gc_generate_from α, + u_l_le := assume ts s hs, generate_open.basic s hs, + choice := λg hg, topological_space.mk_of_closure g + (subset.antisymm hg $ le_generate_from_iff_subset_is_open.1 $ le_rfl), choice_eq := assume s hs, mk_of_closure_sets } -lemma generate_from_mono {α} {g₁ g₂ : set (set α)} (h : g₁ ⊆ g₂) : - topological_space.generate_from g₁ ≤ topological_space.generate_from g₂ := -(gi_generate_from _).gc.monotone_l h +/-- Topologies on `α` form a complete lattice, with `⊥` the discrete topology + and `⊤` the indiscrete topology. The infimum of a collection of topologies + is the topology generated by all their open sets, while the supremum is the + topology whose open sets are those sets open in every member of the collection. -/ +instance : complete_lattice (topological_space α) := +(gci_generate_from α).lift_complete_lattice + +@[mono] lemma generate_from_anti {α} {g₁ g₂ : set (set α)} (h : g₁ ⊆ g₂) : + generate_from g₂ ≤ generate_from g₁ := +(gc_generate_from _).monotone_u h lemma generate_from_set_of_is_open (t : topological_space α) : - topological_space.generate_from {s | t.is_open s} = t := -(gi_generate_from α).l_u_eq t + generate_from {s | is_open[t] s} = t := +(gci_generate_from α).u_l_eq t lemma left_inverse_generate_from : - left_inverse topological_space.generate_from (λ t : topological_space α, {s | t.is_open s}) := -(gi_generate_from α).left_inverse_l_u + left_inverse generate_from (λ t : topological_space α, {s | is_open[t] s}) := +(gci_generate_from α).u_l_left_inverse lemma generate_from_surjective : - surjective (topological_space.generate_from : set (set α) → topological_space α) := -(gi_generate_from α).l_surjective - -lemma set_of_is_open_injective : injective (λ t : topological_space α, {s | t.is_open s}) := -(gi_generate_from α).u_injective + surjective (generate_from : set (set α) → topological_space α) := +(gci_generate_from α).u_surjective -/-- The "temporary" order `tmp_order` on `topological_space α`, i.e. the inclusion order, is a -complete lattice. (Note that later `topological_space α` will equipped with the dual order to -`tmp_order`). -/ -def tmp_complete_lattice {α : Type u} : complete_lattice (topological_space α) := -(gi_generate_from α).lift_complete_lattice +lemma set_of_is_open_injective : injective (λ t : topological_space α, {s | is_open[t] s}) := +(gci_generate_from α).l_injective -instance : has_le (topological_space α) := -{ le := λ t s, s.is_open ≤ t.is_open } - -protected lemma topological_space.le_def {α} {t s : topological_space α} : - t ≤ s ↔ s.is_open ≤ t.is_open := iff.rfl +end lattice -lemma is_open.mono {α} {t₁ t₂ : topological_space α} {s : set α} (hs : @is_open α t₂ s) - (h : t₁ ≤ t₂) : @is_open α t₁ s := h s hs +end topological_space -lemma is_closed.mono {α} {t₁ t₂ : topological_space α} {s : set α} (hs : @is_closed α t₂ s) - (h : t₁ ≤ t₂) : @is_closed α t₁ s := -(@is_open_compl_iff α t₁ s).mp $ hs.is_open_compl.mono h +section lattice -/-- The ordering on topologies on the type `α`. - `t ≤ s` if every set open in `s` is also open in `t` (`t` is finer than `s`). -/ -instance : partial_order (topological_space α) := -{ le_antisymm := assume t s h₁ h₂, topological_space_eq $ le_antisymm h₂ h₁, - le_refl := assume t, le_refl t.is_open, - le_trans := assume a b c h₁ h₂, topological_space.le_def.mpr (le_trans h₂ h₁), - ..topological_space.has_le } +variables {α : Type u} {β : Type v} -lemma le_generate_from_iff_subset_is_open {g : set (set α)} {t : topological_space α} : - t ≤ topological_space.generate_from g ↔ g ⊆ {s | t.is_open s} := -generate_from_le_iff_subset_is_open +lemma is_open.mono {α} {t₁ t₂ : topological_space α} {s : set α} (hs : is_open[t₂] s) + (h : t₁ ≤ t₂) : is_open[t₁] s := h s hs -/-- Topologies on `α` form a complete lattice, with `⊥` the discrete topology - and `⊤` the indiscrete topology. The infimum of a collection of topologies - is the topology generated by all their open sets, while the supremum is the - topology whose open sets are those sets open in every member of the collection. -/ -instance : complete_lattice (topological_space α) := -@order_dual.complete_lattice _ tmp_complete_lattice +lemma is_closed.mono {α} {t₁ t₂ : topological_space α} {s : set α} (hs : is_closed[t₂] s) + (h : t₁ ≤ t₂) : is_closed[t₁] s := +(@is_open_compl_iff α t₁ s).mp $ hs.is_open_compl.mono h lemma is_open_implies_is_open_iff {a b : topological_space α} : - (∀ s, a.is_open s → b.is_open s) ↔ b ≤ a := + (∀ s, is_open[a] s → is_open[b] s) ↔ b ≤ a := iff.rfl /-- The only open sets in the indiscrete topology are the empty set and the whole space. -/ lemma topological_space.is_open_top_iff {α} (U : set α) : - (⊤ : topological_space α).is_open U ↔ U = ∅ ∨ U = univ := + is_open[⊤] U ↔ U = ∅ ∨ U = univ := ⟨λ h, begin induction h with V h _ _ _ _ ih₁ ih₂ _ _ ih, { cases h }, { exact or.inr rfl }, @@ -274,9 +249,7 @@ end, by { rintro (rfl|rfl), exacts [@is_open_empty _ ⊤, @is_open_univ _ ⊤] } class discrete_topology (α : Type*) [t : topological_space α] : Prop := (eq_bot [] : t = ⊥) -@[priority 100] -instance discrete_topology_bot (α : Type*) : @discrete_topology α ⊥ := -{ eq_bot := rfl } +lemma discrete_topology_bot (α : Type*) : @discrete_topology α ⊥ := @discrete_topology.mk α ⊥ rfl @[simp] lemma is_open_discrete [topological_space α] [discrete_topology α] (s : set α) : is_open s := @@ -284,22 +257,16 @@ instance discrete_topology_bot (α : Type*) : @discrete_topology α ⊥ := @[simp] lemma is_closed_discrete [topological_space α] [discrete_topology α] (s : set α) : is_closed s := -is_open_compl_iff.1 $ (discrete_topology.eq_bot α).symm ▸ trivial +is_open_compl_iff.1 $ is_open_discrete _ @[nontriviality] lemma continuous_of_discrete_topology [topological_space α] [discrete_topology α] [topological_space β] {f : α → β} : continuous f := -continuous_def.2 $ λs hs, is_open_discrete _ +continuous_def.2 $ λ s hs, is_open_discrete _ -lemma nhds_bot (α : Type*) : (@nhds α ⊥) = pure := -begin - refine le_antisymm _ (@pure_le_nhds α ⊥), - assume a s hs, - exact @is_open.mem_nhds α ⊥ a s trivial hs -end - -lemma nhds_discrete (α : Type*) [topological_space α] [discrete_topology α] : (@nhds α _) = pure := -(discrete_topology.eq_bot α).symm ▸ nhds_bot α +@[simp] lemma nhds_discrete (α : Type*) [topological_space α] [discrete_topology α] : + (@nhds α _) = pure := +le_antisymm (λ _ s hs, (is_open_discrete s).mem_nhds hs) pure_le_nhds lemma mem_nhds_discrete [topological_space α] [discrete_topology α] {x : α} {s : set α} : s ∈ 𝓝 x ↔ x ∈ s := @@ -307,8 +274,11 @@ by rw [nhds_discrete, mem_pure] lemma le_of_nhds_le_nhds {t₁ t₂ : topological_space α} (h : ∀x, @nhds α t₁ x ≤ @nhds α t₂ x) : t₁ ≤ t₂ := -assume s, show @is_open α t₂ s → @is_open α t₁ s, - by { simp only [is_open_iff_nhds, le_principal_iff], exact assume hs a ha, h _ $ hs _ ha } +begin + intro s, + rw [@is_open_iff_mem_nhds _ t₁, @is_open_iff_mem_nhds α t₂], + exact λ hs a ha, h _ (hs _ ha) +end lemma eq_of_nhds_eq_nhds {t₁ t₂ : topological_space α} (h : ∀x, @nhds α t₁ x = @nhds α t₂ x) : t₁ = t₂ := @@ -316,41 +286,32 @@ le_antisymm (le_of_nhds_le_nhds $ assume x, le_of_eq $ h x) (le_of_nhds_le_nhds $ assume x, le_of_eq $ (h x).symm) -lemma eq_bot_of_singletons_open {t : topological_space α} (h : ∀ x, t.is_open {x}) : t = ⊥ := +lemma eq_bot_of_singletons_open {t : topological_space α} (h : ∀ x, is_open[t] {x}) : t = ⊥ := bot_unique $ λ s hs, bUnion_of_singleton s ▸ is_open_bUnion (λ x _, h x) lemma forall_open_iff_discrete {X : Type*} [topological_space X] : (∀ s : set X, is_open s) ↔ discrete_topology X := -⟨λ h, ⟨by { ext U , show is_open U ↔ true, simp [h U] }⟩, λ a, @is_open_discrete _ _ a⟩ +⟨λ h, ⟨eq_bot_of_singletons_open $ λ _, h _⟩, @is_open_discrete _ _⟩ lemma singletons_open_iff_discrete {X : Type*} [topological_space X] : (∀ a : X, is_open ({a} : set X)) ↔ discrete_topology X := ⟨λ h, ⟨eq_bot_of_singletons_open h⟩, λ a _, @is_open_discrete _ _ a _⟩ +lemma discrete_topology_iff_singleton_mem_nhds [topological_space α] : + discrete_topology α ↔ ∀ x : α, {x} ∈ 𝓝 x := +by simp only [← singletons_open_iff_discrete, is_open_iff_mem_nhds, mem_singleton_iff, forall_eq] + /-- This lemma characterizes discrete topological spaces as those whose singletons are neighbourhoods. -/ lemma discrete_topology_iff_nhds [topological_space α] : discrete_topology α ↔ ∀ x : α, 𝓝 x = pure x := -begin - split ; introI h, - { intro x, - rw nhds_discrete }, - { constructor, - apply eq_of_nhds_eq_nhds, - simp [h, nhds_discrete] }, -end +by simp only [discrete_topology_iff_singleton_mem_nhds, ← nhds_ne_bot.le_pure_iff, + le_pure_iff] lemma discrete_topology_iff_nhds_ne [topological_space α] : discrete_topology α ↔ ∀ x : α, 𝓝[≠] x = ⊥ := -begin - rw discrete_topology_iff_nhds, - apply forall_congr (λ x, _), - rw [nhds_within, inf_principal_eq_bot, compl_compl], - split ; intro h, - { rw h, - exact singleton_mem_pure }, - { exact le_antisymm (le_pure_iff.mpr h) (pure_le_nhds x) } -end +by simp only [discrete_topology_iff_singleton_mem_nhds, nhds_within, inf_principal_eq_bot, + compl_compl] end lattice @@ -362,10 +323,10 @@ variables {α : Type*} {β : Type*} {γ : Type*} makes `f` continuous. -/ def topological_space.induced {α : Type u} {β : Type v} (f : α → β) (t : topological_space β) : topological_space α := -{ is_open := λs, ∃s', t.is_open s' ∧ f ⁻¹' s' = s, - is_open_univ := ⟨univ, t.is_open_univ, preimage_univ⟩, +{ is_open := λs, ∃ s', is_open s' ∧ f ⁻¹' s' = s, + is_open_univ := ⟨univ, is_open_univ, preimage_univ⟩, is_open_inter := by rintro s₁ s₂ ⟨s'₁, hs₁, rfl⟩ ⟨s'₂, hs₂, rfl⟩; - exact ⟨s'₁ ∩ s'₂, t.is_open_inter _ _ hs₁ hs₂, preimage_inter⟩, + exact ⟨s'₁ ∩ s'₂, hs₁.inter hs₂, preimage_inter⟩, is_open_sUnion := assume s h, begin simp only [classical.skolem] at h, @@ -377,15 +338,11 @@ def topological_space.induced {α : Type u} {β : Type v} (f : α → β) (t : t end } lemma is_open_induced_iff [t : topological_space β] {s : set α} {f : α → β} : - @is_open α (t.induced f) s ↔ (∃t, is_open t ∧ f ⁻¹' t = s) := -iff.rfl - -lemma is_open_induced_iff' [t : topological_space β] {s : set α} {f : α → β} : - (t.induced f).is_open s ↔ (∃t, is_open t ∧ f ⁻¹' t = s) := + is_open[t.induced f] s ↔ (∃t, is_open t ∧ f ⁻¹' t = s) := iff.rfl lemma is_closed_induced_iff [t : topological_space β] {s : set α} {f : α → β} : - @is_closed α (t.induced f) s ↔ (∃t, is_closed t ∧ f ⁻¹' t = s) := + 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 compl_surjective.exists.trans (by simp only [preimage_compl, compl_inj_iff]) @@ -396,15 +353,13 @@ end makes `f` continuous. -/ def topological_space.coinduced {α : Type u} {β : Type v} (f : α → β) (t : topological_space α) : topological_space β := -{ is_open := λs, t.is_open (f ⁻¹' s), - is_open_univ := by rw preimage_univ; exact t.is_open_univ, - is_open_inter := assume s₁ s₂ h₁ h₂, by rw preimage_inter; exact t.is_open_inter _ _ h₁ h₂, - is_open_sUnion := assume s h, by rw [preimage_sUnion]; exact (@is_open_Union _ _ t _ $ assume i, - show is_open (⋃ (H : i ∈ s), f ⁻¹' i), from - @is_open_Union _ _ t _ $ assume hi, h i hi) } +{ is_open := λ s, is_open[t] (f ⁻¹' s), + is_open_univ := t.is_open_univ, + is_open_inter := λ _ _ h₁ h₂, h₁.inter h₂, + is_open_sUnion := λ s h, by simpa only [preimage_sUnion] using is_open_bUnion h } lemma is_open_coinduced {t : topological_space α} {s : set β} {f : α → β} : - @is_open β (topological_space.coinduced f t) s ↔ is_open (f ⁻¹' s) := + is_open[t.coinduced f] s ↔ is_open (f ⁻¹' s) := iff.rfl lemma preimage_nhds_coinduced [topological_space α] {π : α → β} {s : set β} @@ -424,9 +379,7 @@ lemma continuous.coinduced_le (h : @continuous α β t t' f) : lemma coinduced_le_iff_le_induced {f : α → β} {tα : topological_space α} {tβ : topological_space β} : tα.coinduced f ≤ tβ ↔ tα ≤ tβ.induced f := -iff.intro - (assume h s ⟨t, ht, hst⟩, hst ▸ h _ ht) - (assume h s hs, show tα.is_open (f ⁻¹' s), from h _ ⟨s, hs, rfl⟩) +⟨λ h s ⟨t, ht, hst⟩, hst ▸ h _ ht, λ h s hs, h _ ⟨s, hs, rfl⟩⟩ lemma continuous.le_induced (h : @continuous α β t t' f) : t ≤ t'.induced f := @@ -489,11 +442,8 @@ begin ext t U, split, { rintros ⟨V, hV, rfl⟩, - change t.is_open (e ⁻¹' _), - rwa [← preimage_comp, ← equiv.coe_trans, equiv.self_trans_symm] }, - { intros hU, - refine ⟨e ⁻¹' U, hU, _⟩, - rw [← preimage_comp, ← equiv.coe_trans, equiv.symm_trans_self, equiv.coe_refl, preimage_id] } + rwa [is_open_coinduced, e.preimage_symm_preimage] }, + { exact λ hU, ⟨e ⁻¹' U, hU, e.symm_preimage_preimage _⟩ } end lemma equiv.coinduced_symm {α β : Type*} (e : α ≃ β) : @@ -509,7 +459,7 @@ open topological_space variables {α : Type u} {β : Type v} instance inhabited_topological_space {α : Type u} : inhabited (topological_space α) := -⟨⊤⟩ +⟨⊥⟩ @[priority 100] instance subsingleton.unique_topological_space [subsingleton α] : @@ -536,6 +486,9 @@ instance : discrete_topology ℕ := ⟨rfl⟩ instance : topological_space ℤ := ⊥ instance : discrete_topology ℤ := ⟨rfl⟩ +instance {n} : topological_space (fin n) := ⊥ +instance {n} : discrete_topology (fin n) := ⟨rfl⟩ + instance sierpinski_space : topological_space Prop := generate_from {{true}} @@ -608,7 +561,7 @@ begin end lemma is_open_singleton_nhds_adjoint {α : Type*} {a b : α} (f : filter α) (hb : b ≠ a) : - @is_open α (nhds_adjoint a f) {b} := + is_open[nhds_adjoint a f] {b} := begin rw is_open_singleton_iff_nhds_eq_pure, exact nhds_adjoint_nhds_of_ne a f hb @@ -634,7 +587,7 @@ begin end lemma le_nhds_adjoint_iff {α : Type*} (a : α) (f : filter α) (t : topological_space α) : - t ≤ nhds_adjoint a f ↔ (@nhds α t a ≤ pure a ⊔ f ∧ ∀ b, b ≠ a → t.is_open {b}) := + t ≤ nhds_adjoint a f ↔ (@nhds α t a ≤ pure a ⊔ f ∧ ∀ b, b ≠ a → is_open[t] {b}) := begin change _ ↔ _ ∧ ∀ (b : α), b ≠ a → is_open {b}, rw [le_nhds_adjoint_iff', and.congr_right_iff], @@ -654,7 +607,7 @@ lemma nhds_inf {t₁ t₂ : topological_space α} {a : α} : lemma nhds_top {a : α} : @nhds α ⊤ a = ⊤ := (gc_nhds a).u_top lemma is_open_sup {t₁ t₂ : topological_space α} {s : set α} : - @is_open α (t₁ ⊔ t₂) s ↔ @is_open α t₁ s ∧ @is_open α t₂ s := + is_open[t₁ ⊔ t₂] s ↔ is_open[t₁] s ∧ is_open[t₂] s := iff.rfl local notation `cont` := @continuous _ _ @@ -808,10 +761,10 @@ variables {α : Type*} {β : Type*} variables [t : topological_space β] {f : α → β} theorem is_open_induced_eq {s : set α} : - @is_open _ (induced f t) s ↔ s ∈ preimage f '' {s | is_open s} := + is_open[induced f t] s ↔ s ∈ preimage f '' {s | is_open s} := iff.rfl -theorem is_open_induced {s : set β} (h : is_open s) : (induced f t).is_open (f ⁻¹' s) := +theorem is_open_induced {s : set β} (h : is_open s) : is_open[induced f t] (f ⁻¹' s) := ⟨s, h, rfl⟩ lemma map_nhds_induced_eq (a : α) : map f (@nhds α (induced f t) a) = 𝓝[range f] (f a) := @@ -826,7 +779,7 @@ lemma closure_induced [t : topological_space β] {f : α → β} {a : α} {s : 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 := + 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 @@ -863,62 +816,56 @@ 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 _ (topological_space α)ᵒᵈ a₁ a₂ _ _ _ _ - (λ g t, generate_from_le_iff_subset_is_open) +(topological_space.gc_generate_from α).u_inf 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 _ (topological_space α)ᵒᵈ t₁ t₂ _ _ _ _ - (λ g t, generate_from_le_iff_subset_is_open) + {s | is_open[t₁ ⊔ t₂] s} = {s | is_open[t₁] s} ∩ {s | is_open[t₂] s} := +rfl 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 _ (topological_space α)ᵒᵈ _ _ _ _ _ - (λ g t, generate_from_le_iff_subset_is_open) f +(topological_space.gc_generate_from α).u_infi 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 _ (topological_space α)ᵒᵈ _ _ _ _ _ - (λ g t, generate_from_le_iff_subset_is_open) t + {s | is_open[⨆ i, t i] s} = ⋂ i, {s | is_open[t i] s} := +(topological_space.gc_generate_from α).l_supr lemma generate_from_sUnion {S : set (set (set α))} : topological_space.generate_from (⋃₀ S) = (⨅ s ∈ S, topological_space.generate_from s) := -@galois_connection.l_Sup _ (topological_space α)ᵒᵈ _ _ _ _ - (λ g t, generate_from_le_iff_subset_is_open) S +(topological_space.gc_generate_from α).u_Inf 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 _ (topological_space α)ᵒᵈ _ _ _ _ - (λ g t, generate_from_le_iff_subset_is_open) T + {s | is_open[Sup T] s} = ⋂ t ∈ T, {s | is_open[t] s} := +(topological_space.gc_generate_from α).l_Sup 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 _ (topological_space α)ᵒᵈ _ _ _ _ (gi_generate_from α) a b + topological_space.generate_from ({s | is_open[a] s} ∪ {s | is_open[b] s}) = a ⊓ b := +(topological_space.gci_generate_from α).u_inf_l 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 _ (topological_space α)ᵒᵈ _ _ _ _ (gi_generate_from α) _ f + topological_space.generate_from (⋃ i, {s | is_open[f i] s}) = ⨅ i, (f i) := +(topological_space.gci_generate_from α).u_infi_l 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 _ (topological_space α)ᵒᵈ _ _ _ _ (gi_generate_from α) a b + topological_space.generate_from ({s | is_open[a] s} ∩ {s | is_open[b] s}) = a ⊔ b := +(topological_space.gci_generate_from α).u_sup_l 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 _ (topological_space α)ᵒᵈ _ _ _ _ (gi_generate_from α) _ f + topological_space.generate_from (⋂ i, {s | is_open[f i] s}) = ⨆ i, (f i) := +(topological_space.gci_generate_from α).u_supr_l 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) : + (hf : ∀ i, {s | is_open[topological_space.generate_from (f i)] 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 _ (topological_space α)ᵒᵈ _ _ _ _ (gi_generate_from α) _ f hf +(topological_space.gci_generate_from α).u_supr_of_lu_eq_self f hf variables {t : ι → topological_space α} -lemma is_open_supr_iff {s : set α} : @is_open _ (⨆ i, t i) s ↔ ∀ i, @is_open _ (t i) s := -show s ∈ set_of (supr t).is_open ↔ s ∈ {x : set α | ∀ (i : ι), (t i).is_open x}, +lemma is_open_supr_iff {s : set α} : is_open[⨆ i, t i] s ↔ ∀ i, is_open[t i] s := +show s ∈ set_of (is_open[supr t]) ↔ s ∈ {x : set α | ∀ (i : ι), is_open[t i] x}, by simp [set_of_is_open_supr] -lemma is_closed_supr_iff {s : set α} : @is_closed _ (⨆ i, t i) s ↔ ∀ i, @is_closed _ (t i) s := +lemma is_closed_supr_iff {s : set α} : is_closed[⨆ i, t i] s ↔ ∀ i, is_closed[t i] s := by simp [← is_open_compl_iff, is_open_supr_iff] end infi diff --git a/src/topology/separation.lean b/src/topology/separation.lean index a1f19d14f4381..1dafe29e01bc0 100644 --- a/src/topology/separation.lean +++ b/src/topology/separation.lean @@ -770,34 +770,6 @@ begin rw ← induced_compose, end -/-- The topology pulled-back under an inclusion `f : X → Y` from the discrete topology (`⊥`) is the -discrete topology. -This version does not assume the choice of a topology on either the source `X` -nor the target `Y` of the inclusion `f`. -/ -lemma induced_bot {X Y : Type*} {f : X → Y} (hf : function.injective f) : - topological_space.induced f ⊥ = ⊥ := -eq_of_nhds_eq_nhds (by simp [nhds_induced, ← set.image_singleton, hf.preimage_image, nhds_bot]) - -/-- The topology induced under an inclusion `f : X → Y` from the discrete topological space `Y` -is the discrete topology on `X`. -/ -lemma discrete_topology_induced {X Y : Type*} [tY : topological_space Y] [discrete_topology Y] - {f : X → Y} (hf : function.injective f) : @discrete_topology X (topological_space.induced f tY) := -by apply discrete_topology.mk; by rw [discrete_topology.eq_bot Y, induced_bot hf] - -lemma embedding.discrete_topology {X Y : Type*} [topological_space X] [tY : topological_space Y] - [discrete_topology Y] {f : X → Y} (hf : embedding f) : discrete_topology X := -⟨by rw [hf.induced, discrete_topology.eq_bot Y, induced_bot hf.inj]⟩ - -/-- Let `s, t ⊆ X` be two subsets of a topological space `X`. If `t ⊆ s` and the topology induced -by `X`on `s` is discrete, then also the topology induces on `t` is discrete. -/ -lemma discrete_topology.of_subset {X : Type*} [topological_space X] {s t : set X} - (ds : discrete_topology s) (ts : t ⊆ s) : - discrete_topology t := -begin - rw [topological_space.subset_trans ts, ds.eq_bot], - exact {eq_bot := induced_bot (set.inclusion_injective ts)} -end - /-- A T₂ space, also known as a Hausdorff space, is one in which for every `x ≠ y` there exists disjoint open sets around `x` and `y`. This is the most widely used of the separation axioms. -/ @@ -1459,7 +1431,7 @@ begin letI := Inf T, have : ∀ a, (𝓝 a).has_basis (λ If : Σ I : set T, I → set X, - If.1.finite ∧ ∀ i : If.1, If.2 i ∈ @nhds X i a ∧ @is_closed X i (If.2 i)) + If.1.finite ∧ ∀ i : If.1, If.2 i ∈ @nhds X i a ∧ is_closed[↑i] (If.2 i)) (λ If, ⋂ i : If.1, If.snd i), { intro a, rw [nhds_Inf, ← infi_subtype''], @@ -1722,7 +1694,7 @@ begin -- We do this by showing that any disjoint cover by two closed sets implies -- that one of these closed sets must contain our whole thing. -- To reduce to the case where the cover is disjoint on all of `α` we need that `s` is closed - have hs : @is_closed α _ (⋂ (Z : {Z : set α // is_clopen Z ∧ x ∈ Z}), Z) := + have hs : is_closed (⋂ (Z : {Z : set α // is_clopen Z ∧ x ∈ Z}), Z : set α) := is_closed_Inter (λ Z, Z.2.1.2), rw (is_preconnected_iff_subset_of_fully_disjoint_closed hs), intros a b ha hb hab ab_disj, diff --git a/src/topology/subset_properties.lean b/src/topology/subset_properties.lean index c255c89a188b4..0d23021b619db 100644 --- a/src/topology/subset_properties.lean +++ b/src/topology/subset_properties.lean @@ -571,7 +571,7 @@ lemma tendsto.is_compact_insert_range_of_cofinite {f : ι → α} {a} (hf : tendsto f cofinite (𝓝 a)) : is_compact (insert a (range f)) := begin - letI : topological_space ι := ⊥, haveI : discrete_topology ι := ⟨rfl⟩, + letI : topological_space ι := ⊥, haveI := discrete_topology_bot ι, rw ← cocompact_eq_cofinite at hf, exact hf.is_compact_insert_range_of_cocompact continuous_of_discrete_topology end diff --git a/src/topology/uniform_space/basic.lean b/src/topology/uniform_space/basic.lean index aa67e33aabecb..c8a3c0692295d 100644 --- a/src/topology/uniform_space/basic.lean +++ b/src/topology/uniform_space/basic.lean @@ -256,12 +256,13 @@ lemma uniform_space.core_eq : A metric space has a natural uniformity, and a uniform space has a natural topology. A topological group also has a natural uniformity, even when it is not metrizable. -/ class uniform_space (α : Type u) extends topological_space α, uniform_space.core α := -(is_open_uniformity : ∀s, is_open s ↔ (∀x∈s, { p : α × α | p.1 = x → p.2 ∈ s } ∈ uniformity)) +(is_open_uniformity : ∀s, @_root_.is_open _ to_topological_space s ↔ + (∀x∈s, { p : α × α | p.1 = x → p.2 ∈ s } ∈ uniformity)) /-- Alternative constructor for `uniform_space α` when a topology is already given. -/ @[pattern] def uniform_space.mk' {α} (t : topological_space α) (c : uniform_space.core α) - (is_open_uniformity : ∀s:set α, t.is_open s ↔ + (is_open_uniformity : ∀s:set α, is_open s ↔ (∀x∈s, { p : α × α | p.1 = x → p.2 ∈ s } ∈ c.uniformity)) : uniform_space α := ⟨c, is_open_uniformity⟩ @@ -281,11 +282,17 @@ def uniform_space.of_core_eq {α : Type u} (u : uniform_space.core α) (t : topo lemma uniform_space.to_core_to_topological_space (u : uniform_space α) : u.to_core.to_topological_space = u.to_topological_space := -topological_space_eq $ funext $ assume s, - by rw [uniform_space.core.to_topological_space, uniform_space.is_open_uniformity] +topological_space_eq $ funext $ λ s, by rw [uniform_space.is_open_uniformity, is_open_mk] + +/-- The uniformity is a filter on α × α (inferred from an ambient uniform space + structure on α). -/ +def uniformity (α : Type u) [uniform_space α] : filter (α × α) := + (@uniform_space.to_core α _).uniformity + +localized "notation (name := uniformity_of) `𝓤[` u `]` := @uniformity hole! u" in topology @[ext] -lemma uniform_space_eq : ∀{u₁ u₂ : uniform_space α}, u₁.uniformity = u₂.uniformity → u₁ = u₂ +lemma uniform_space_eq : ∀ {u₁ u₂ : uniform_space α}, 𝓤[u₁] = 𝓤[u₂] → u₁ = u₂ | (uniform_space.mk' t₁ u₁ o₁) (uniform_space.mk' t₂ u₂ o₂) h := have u₁ = u₂, from uniform_space.core_eq h, have t₁ = t₂, from topological_space_eq $ funext $ assume s, by rw [o₁, o₂]; simp [this], @@ -309,11 +316,6 @@ u.of_core_eq_to_core _ _ section uniform_space variables [uniform_space α] -/-- The uniformity is a filter on α × α (inferred from an ambient uniform space - structure on α). -/ -def uniformity (α : Type u) [uniform_space α] : filter (α × α) := - (@uniform_space.to_core α _).uniformity - localized "notation (name := uniformity) `𝓤` := uniformity" in uniformity lemma is_open_uniformity {s : set α} : @@ -1013,7 +1015,7 @@ instance : partial_order (uniform_space α) := instance : has_Inf (uniform_space α) := ⟨assume s, uniform_space.of_core -{ uniformity := (⨅u∈s, @uniformity α u), +{ uniformity := (⨅u∈s, 𝓤[u]), refl := le_infi $ assume u, le_infi $ assume hu, u.refl, symm := le_infi $ assume u, le_infi $ assume hu, le_trans (map_mono $ infi_le_of_le _ $ infi_le _ hu) u.symm, @@ -1022,13 +1024,11 @@ instance : has_Inf (uniform_space α) := private lemma Inf_le {tt : set (uniform_space α)} {t : uniform_space α} (h : t ∈ tt) : Inf tt ≤ t := -show (⨅u∈tt, @uniformity α u) ≤ t.uniformity, - from infi_le_of_le t $ infi_le _ h +show (⨅ u ∈ tt, 𝓤[u]) ≤ 𝓤[t], from infi₂_le t h private lemma le_Inf {tt : set (uniform_space α)} {t : uniform_space α} (h : ∀t'∈tt, t ≤ t') : t ≤ Inf tt := -show t.uniformity ≤ (⨅u∈tt, @uniformity α u), - from le_infi $ assume t', le_infi $ assume ht', h t' ht' +show 𝓤[t] ≤ (⨅ u ∈ tt, 𝓤[u]), from le_infi₂ h instance : has_top (uniform_space α) := ⟨uniform_space.of_core { uniformity := ⊤, refl := le_top, symm := le_top, comp := le_top }⟩ @@ -1074,24 +1074,10 @@ instance : complete_lattice (uniform_space α) := Inf_le := λ s a ha, Inf_le ha, ..uniform_space.partial_order } -lemma infi_uniformity {ι : Sort*} {u : ι → uniform_space α} : - (infi u).uniformity = (⨅i, (u i).uniformity) := -show (⨅a (h : ∃i:ι, u i = a), a.uniformity) = _, from -le_antisymm - (le_infi $ assume i, infi_le_of_le (u i) $ infi_le _ ⟨i, rfl⟩) - (le_infi $ assume a, le_infi $ assume ⟨i, (ha : u i = a)⟩, ha ▸ infi_le _ _) - -lemma infi_uniformity' {ι : Sort*} {u : ι → uniform_space α} : - @uniformity α (infi u) = (⨅i, @uniformity α (u i)) := -infi_uniformity +lemma infi_uniformity {ι : Sort*} {u : ι → uniform_space α} : 𝓤[infi u] = (⨅i, 𝓤[u i]) := +infi_range -lemma inf_uniformity {u v : uniform_space α} : - (u ⊓ v).uniformity = u.uniformity ⊓ v.uniformity := -rfl - -lemma inf_uniformity' {u v : uniform_space α} : - @uniformity α (u ⊓ v) = @uniformity α u ⊓ @uniformity α v := -rfl +lemma inf_uniformity {u v : uniform_space α} : 𝓤[u ⊓ v] = 𝓤[u] ⊓ 𝓤[v] := rfl instance inhabited_uniform_space : inhabited (uniform_space α) := ⟨⊥⟩ instance inhabited_uniform_space_core : inhabited (uniform_space.core α) := @@ -1100,7 +1086,7 @@ instance inhabited_uniform_space_core : inhabited (uniform_space.core α) := /-- Given `f : α → β` and a uniformity `u` on `β`, the inverse image of `u` under `f` is the inverse image in the filter sense of the induced function `α × α → β × β`. -/ def uniform_space.comap (f : α → β) (u : uniform_space β) : uniform_space α := -{ uniformity := u.uniformity.comap (λp:α×α, (f p.1, f p.2)), +{ uniformity := 𝓤[u].comap (λp:α×α, (f p.1, f p.2)), to_topological_space := u.to_topological_space.induced f, refl := le_trans (by simp; exact assume ⟨a, b⟩ (h : a = b), h ▸ rfl) (comap_mono u.refl), symm := by simp [tendsto_comap_iff, prod.swap, (∘)]; @@ -1115,17 +1101,16 @@ def uniform_space.comap (f : α → β) (u : uniform_space β) : uniform_space is_open_uniformity := λ s, by simp only [is_open_fold, is_open_induced, is_open_iff_mem_nhds, nhds_induced, nhds_eq_comap_uniformity, comap_comap, ← mem_comap_prod_mk, ← uniformity] } -lemma uniformity_comap [uniform_space α] [uniform_space β] {f : α → β} - (h : ‹uniform_space α› = uniform_space.comap f ‹uniform_space β›) : - 𝓤 α = comap (prod.map f f) (𝓤 β) := -by { rw h, refl } +lemma uniformity_comap [uniform_space β] (f : α → β) : + 𝓤[uniform_space.comap f ‹_›] = comap (prod.map f f) (𝓤 β) := +rfl -lemma uniform_space_comap_id {α : Type*} : uniform_space.comap (id : α → α) = id := -by ext u ; dsimp only [uniform_space.comap, id] ; rw [prod.id_prod, filter.comap_id] +@[simp] lemma uniform_space_comap_id {α : Type*} : uniform_space.comap (id : α → α) = id := +by { ext : 2, rw [uniformity_comap, prod.map_id, comap_id] } lemma uniform_space.comap_comap {α β γ} [uγ : uniform_space γ] {f : α → β} {g : β → γ} : uniform_space.comap (g ∘ f) uγ = uniform_space.comap f (uniform_space.comap g uγ) := -by ext ; dsimp only [uniform_space.comap] ; rw filter.comap_comap +by { ext1, simp only [uniformity_comap, comap_comap, prod.map_comp_map] } lemma uniform_space.comap_inf {α γ} {u₁ u₂ : uniform_space γ} {f : α → γ} : (u₁ ⊓ u₂).comap f = u₁.comap f ⊓ u₂.comap f := @@ -1135,8 +1120,7 @@ lemma uniform_space.comap_infi {ι α γ} {u : ι → uniform_space γ} {f : α (⨅ i, u i).comap f = ⨅ i, (u i).comap f := begin ext : 1, - change (𝓤 _) = (𝓤 _), - simp [uniformity_comap rfl, infi_uniformity'] + simp [uniformity_comap, infi_uniformity] end lemma uniform_space.comap_mono {α γ} {f : α → γ} : @@ -1144,7 +1128,7 @@ lemma uniform_space.comap_mono {α γ} {f : α → γ} : begin intros u₁ u₂ hu, change (𝓤 _) ≤ (𝓤 _), - rw uniformity_comap rfl, + rw uniformity_comap, exact comap_mono hu end @@ -1194,11 +1178,8 @@ lemma to_topological_space_infi {ι : Sort*} {u : ι → uniform_space α} : (infi u).to_topological_space = ⨅i, (u i).to_topological_space := begin refine (eq_of_nhds_eq_nhds $ assume a, _), - rw [nhds_infi, nhds_eq_uniformity], - change (infi u).uniformity.lift' (preimage $ prod.mk a) = _, - rw [infi_uniformity, lift'_infi_of_map_univ _ preimage_univ], - { simp only [nhds_eq_uniformity], refl }, - { exact ball_inter _ } + simp only [nhds_infi, nhds_eq_uniformity, infi_uniformity], + exact lift'_infi_of_map_univ (ball_inter _) preimage_univ end lemma to_topological_space_Inf {s : set (uniform_space α)} : @@ -1235,27 +1216,27 @@ lemma uniform_continuous_Inf_dom {f : α → β} {u₁ : set (uniform_space α)} {u : uniform_space α} (h₁ : u ∈ u₁) (hf : @@uniform_continuous u u₂ f) : @@uniform_continuous (Inf u₁) u₂ f := begin - rw [uniform_continuous, Inf_eq_infi', infi_uniformity'], + rw [uniform_continuous, Inf_eq_infi', infi_uniformity], exact tendsto_infi' ⟨u, h₁⟩ hf end lemma uniform_continuous_Inf_rng {f : α → β} {u₁ : uniform_space α} {u₂ : set (uniform_space β)} (h : ∀u∈u₂, @@uniform_continuous u₁ u f) : @@uniform_continuous u₁ (Inf u₂) f := begin - rw [uniform_continuous, Inf_eq_infi', infi_uniformity'], + rw [uniform_continuous, Inf_eq_infi', infi_uniformity], exact tendsto_infi.mpr (λ ⟨u, hu⟩, h u hu) end lemma uniform_continuous_infi_dom {f : α → β} {u₁ : ι → uniform_space α} {u₂ : uniform_space β} {i : ι} (hf : @@uniform_continuous (u₁ i) u₂ f) : @@uniform_continuous (infi u₁) u₂ f := begin - rw [uniform_continuous, infi_uniformity'], + rw [uniform_continuous, infi_uniformity], exact tendsto_infi' i hf end lemma uniform_continuous_infi_rng {f : α → β} {u₁ : uniform_space α} {u₂ : ι → uniform_space β} (h : ∀i, @@uniform_continuous u₁ (u₂ i) f) : @@uniform_continuous u₁ (infi u₂) f := -by rwa [uniform_continuous, infi_uniformity', tendsto_infi] +by rwa [uniform_continuous, infi_uniformity, tendsto_infi] end uniform_continuous_infi @@ -1407,7 +1388,7 @@ end lemma mem_uniform_prod [t₁ : uniform_space α] [t₂ : uniform_space β] {a : set (α × α)} {b : set (β × β)} (ha : a ∈ 𝓤 α) (hb : b ∈ 𝓤 β) : - {p:(α×β)×(α×β) | (p.1.1, p.2.1) ∈ a ∧ (p.1.2, p.2.2) ∈ b } ∈ (@uniformity (α × β) _) := + {p:(α×β)×(α×β) | (p.1.1, p.2.1) ∈ a ∧ (p.1.2, p.2.2) ∈ b } ∈ 𝓤 (α × β) := by rw [uniformity_prod]; exact inter_mem_inf (preimage_mem_comap ha) (preimage_mem_comap hb) lemma tendsto_prod_uniformity_fst [uniform_space α] [uniform_space β] : diff --git a/src/topology/uniform_space/compact_convergence.lean b/src/topology/uniform_space/compact_convergence.lean index 33f714537b1a1..044ef05e4be37 100644 --- a/src/topology/uniform_space/compact_convergence.lean +++ b/src/topology/uniform_space/compact_convergence.lean @@ -264,8 +264,8 @@ begin haveI := hι, exact ⟨⋂ i, compact_open.gen (C i) (U i), h₂.trans hXf, is_open_Inter (λ i, continuous_map.is_open_gen (hC i) (hU i)), h₁⟩, }, - { simp only [le_generate_from_iff_subset_is_open, and_imp, exists_prop, forall_exists_index, - set_of_subset_set_of], + { simp only [topological_space.le_generate_from_iff_subset_is_open, and_imp, exists_prop, + forall_exists_index, set_of_subset_set_of], rintros - K hK U hU rfl f hf, obtain ⟨V, hV, hV', hVf⟩ := compact_conv_nhd_subset_compact_open f hK hU hf, exact filter.mem_of_superset (filter_basis.mem_filter_of_mem _ ⟨⟨K, V⟩, ⟨hK, hV⟩, rfl⟩) hVf, }, diff --git a/src/topology/uniform_space/separation.lean b/src/topology/uniform_space/separation.lean index 1c016b1d08084..5f9cdf39ce5d2 100644 --- a/src/topology/uniform_space/separation.lean +++ b/src/topology/uniform_space/separation.lean @@ -158,8 +158,9 @@ lemma separation_rel_comap {f : α → β} (h : ‹uniform_space α› = uniform_space.comap f ‹uniform_space β›) : 𝓢 α = (prod.map f f) ⁻¹' 𝓢 β := begin + unfreezingI { subst h }, dsimp [separation_rel], - simp_rw [uniformity_comap h, (filter.comap_has_basis (prod.map f f) (𝓤 β)).sInter_sets, + simp_rw [uniformity_comap, (filter.comap_has_basis (prod.map f f) (𝓤 β)).sInter_sets, ← preimage_Inter, sInter_eq_bInter], refl, end @@ -276,7 +277,8 @@ instance separation_setoid.uniform_space {α : Type u} [u : uniform_space α] : u.uniformity.sets_of_superset ht $ assume ⟨a₁, a₂⟩ h₁ h₂, hts (ht' $ setoid.symm h₂) h₁, assume h, u.uniformity.sets_of_superset h $ by simp {contextual := tt}⟩, begin - simp [topological_space.coinduced, u.is_open_uniformity, uniformity, forall_quotient_iff], + simp only [is_open_coinduced, is_open_uniformity, uniformity, forall_quotient_iff, + mem_preimage, mem_map, preimage_set_of_eq, quotient.eq], exact ⟨λh a ha, (this a ha).mp $ h a ha, λh a ha, (this a ha).mpr $ h a ha⟩ end } diff --git a/src/topology/uniform_space/uniform_convergence_topology.lean b/src/topology/uniform_space/uniform_convergence_topology.lean index 51e5a7e1eeac8..4db6b9b24b97a 100644 --- a/src/topology/uniform_space/uniform_convergence_topology.lean +++ b/src/topology/uniform_space/uniform_convergence_topology.lean @@ -320,9 +320,8 @@ begin -- This follows directly from the fact that the upper adjoint in a Galois connection maps -- infimas to infimas. ext : 1, - change uniform_fun.filter α γ (@uniformity _ (⨅ i, u i)) = - @uniformity _ (⨅ i, (𝒰(α, γ, u i))), - rw [infi_uniformity', infi_uniformity'], + change uniform_fun.filter α γ (𝓤[⨅ i, u i]) = 𝓤[⨅ i, 𝒰(α, γ, u i)], + rw [infi_uniformity, infi_uniformity], exact (uniform_fun.gc α γ).u_infi end @@ -391,7 +390,7 @@ begin split, replace hf : (𝓤 β).comap (prod.map f f) = _ := hf.comap_uniformity, change comap (prod.map (of_fun ∘ (∘) f ∘ to_fun) (of_fun ∘ (∘) f ∘ to_fun)) _ = _, - rw [← uniformity_comap rfl] at ⊢ hf, + rw [← uniformity_comap] at ⊢ hf, congr, rw [← uniform_space_eq hf, uniform_fun.comap_eq], refl @@ -477,7 +476,7 @@ begin split, change comap (prod.map (equiv.arrow_prod_equiv_prod_arrow _ _ _) (equiv.arrow_prod_equiv_prod_arrow _ _ _)) _ = _, - rw ← uniformity_comap rfl, + rw ← uniformity_comap, congr, rw [prod.uniform_space, prod.uniform_space, uniform_space.comap_inf, uniform_fun.inf_eq], congr; @@ -502,7 +501,7 @@ protected def uniform_equiv_Pi_comm : uniform_equiv (α →ᵤ Π i, δ i) (Π i begin split, change comap (prod.map function.swap function.swap) _ = _, - rw ← uniformity_comap rfl, + rw ← uniformity_comap, congr, rw [Pi.uniform_space, uniform_space.of_core_eq_to_core, Pi.uniform_space, uniform_space.of_core_eq_to_core, uniform_space.comap_infi, uniform_fun.infi_eq], @@ -593,7 +592,7 @@ protected lemma has_basis_uniformity_of_basis_aux₁ {p : ι → Prop} {s : ι (@uniformity (α →ᵤ[𝔖] β) ((uniform_fun.uniform_space S β).comap S.restrict)).has_basis p (λ i, uniform_on_fun.gen 𝔖 S (s i)) := begin - simp_rw [uniform_on_fun.gen_eq_preimage_restrict, uniformity_comap rfl], + simp_rw [uniform_on_fun.gen_eq_preimage_restrict, uniformity_comap], exact (uniform_fun.has_basis_uniformity_of_basis S β hb).comap _ end @@ -615,7 +614,7 @@ protected lemma has_basis_uniformity_of_basis (h : 𝔖.nonempty) (h' : directed (λ Si : set α × ι, Si.1 ∈ 𝔖 ∧ p Si.2) (λ Si, uniform_on_fun.gen 𝔖 Si.1 (s Si.2)) := begin - simp only [infi_uniformity'], + simp only [infi_uniformity], exact has_basis_binfi_of_directed h (λ S, (uniform_on_fun.gen 𝔖 S) ∘ s) _ (λ S hS, uniform_on_fun.has_basis_uniformity_of_basis_aux₁ α β 𝔖 hb S) (uniform_on_fun.has_basis_uniformity_of_basis_aux₂ α β 𝔖 h' hb) @@ -656,10 +655,8 @@ protected lemma uniform_continuous_restrict (h : s ∈ 𝔖) : uniform_continuous (uniform_fun.of_fun ∘ (s.restrict : (α → β) → (s → β)) ∘ (to_fun 𝔖)) := begin change _ ≤ _, - rw [uniform_on_fun.uniform_space, map_le_iff_le_comap, uniformity, infi_uniformity], - refine infi_le_of_le s _, - rw infi_uniformity, - exact infi_le _ h, + simp only [uniform_on_fun.uniform_space, map_le_iff_le_comap, infi_uniformity], + exact infi₂_le s h end variables {α} @@ -750,7 +747,7 @@ begin split, replace hf : (𝓤 β).comap (prod.map f f) = _ := hf.comap_uniformity, change comap (prod.map (of_fun 𝔖 ∘ (∘) f ∘ to_fun 𝔖) (of_fun 𝔖 ∘ (∘) f ∘ to_fun 𝔖)) _ = _, - rw [← uniformity_comap rfl] at ⊢ hf, + rw [← uniformity_comap] at ⊢ hf, congr, rw [← uniform_space_eq hf, uniform_on_fun.comap_eq], refl @@ -866,17 +863,13 @@ protected def uniform_equiv_prod_arrow [uniform_space γ] : -- that some square commutes. -- We could also deduce this from `uniform_convergence.uniform_equiv_prod_arrow`, but it turns out -- to be more annoying. -(equiv.arrow_prod_equiv_prod_arrow _ _ _).to_uniform_equiv_of_uniform_inducing +((uniform_on_fun.of_fun 𝔖).symm.trans $ (equiv.arrow_prod_equiv_prod_arrow _ _ _).trans $ + (uniform_on_fun.of_fun 𝔖).prod_congr (uniform_on_fun.of_fun 𝔖)) + .to_uniform_equiv_of_uniform_inducing begin split, - change comap (prod.map (equiv.arrow_prod_equiv_prod_arrow _ _ _) - (equiv.arrow_prod_equiv_prod_arrow _ _ _)) _ = _, - rw ← uniformity_comap rfl, - congr, - rw [prod.uniform_space, prod.uniform_space, uniform_space.comap_inf, - uniform_on_fun.inf_eq], - congr; - rw [← uniform_space.comap_comap, uniform_on_fun.comap_eq]; + rw [uniformity_prod, comap_inf, comap_comap, comap_comap, uniform_on_fun.inf_eq, inf_uniformity, + uniform_on_fun.comap_eq, uniform_on_fun.comap_eq, uniformity_comap, uniformity_comap], refl -- the relevant diagram commutes by definition end @@ -897,7 +890,7 @@ protected def uniform_equiv_Pi_comm : begin split, change comap (prod.map function.swap function.swap) _ = _, - rw ← uniformity_comap rfl, + rw ← uniformity_comap, congr, rw [Pi.uniform_space, uniform_space.of_core_eq_to_core, Pi.uniform_space, uniform_space.of_core_eq_to_core, uniform_space.comap_infi, uniform_on_fun.infi_eq], diff --git a/src/topology/uniform_space/uniform_embedding.lean b/src/topology/uniform_space/uniform_embedding.lean index b6958a215268f..261ab5583f2f6 100644 --- a/src/topology/uniform_space/uniform_embedding.lean +++ b/src/topology/uniform_space/uniform_embedding.lean @@ -178,7 +178,8 @@ lemma uniform_embedding_of_spaced_out {α} {f : α → β} {s : set (β × β)} (hf : pairwise (λ x y, (f x, f y) ∉ s)) : @uniform_embedding α β ⊥ ‹_› f := begin - letI : uniform_space α := ⊥, haveI : separated_space α := separated_iff_t2.2 infer_instance, + letI : uniform_space α := ⊥, haveI := discrete_topology_bot α, + haveI : separated_space α := separated_iff_t2.2 infer_instance, exact uniform_inducing.uniform_embedding ⟨comap_uniformity_of_spaced_out hs hf⟩ end From ed60ee25ed00d7a62a0d1e5808092e1324cee451 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Tue, 31 Jan 2023 07:32:42 +0000 Subject: [PATCH 0385/1291] feat(data/{bool,set}): add 3 lemmas (#18332) Add `bool.compl_singleton`, `set.range_inl`, and `set.range_inr`. --- src/data/bool/set.lean | 3 +++ src/data/set/image.lean | 3 +++ 2 files changed, 6 insertions(+) diff --git a/src/data/bool/set.lean b/src/data/bool/set.lean index 87ccdb45690e8..bc466d00bb3ee 100644 --- a/src/data/bool/set.lean +++ b/src/data/bool/set.lean @@ -25,4 +25,7 @@ namespace bool @[simp] lemma range_eq {α : Type*} (f : bool → α) : range f = {f ff, f tt} := by rw [← image_univ, univ_eq, image_pair] +@[simp] lemma compl_singleton (b : bool) : ({b}ᶜ : set bool) = { !b } := +ext $ λ _, eq_bnot_iff.symm + end bool diff --git a/src/data/set/image.lean b/src/data/set/image.lean index be5dcf78c6a91..0b70ec820d9d6 100644 --- a/src/data/set/image.lean +++ b/src/data/set/image.lean @@ -640,6 +640,9 @@ prod.snd_surjective.range_eq range (eval i : (Π i, α i) → α i) = univ := (surjective_eval i).range_eq +theorem range_inl : range (@sum.inl α β) = {x | x.is_left} := by ext (_|_); simp +theorem range_inr : range (@sum.inr α β) = {x | x.is_right} := by ext (_|_); simp + theorem is_compl_range_inl_range_inr : is_compl (range $ @sum.inl α β) (range sum.inr) := is_compl.of_le (by { rintro y ⟨⟨x₁, rfl⟩, ⟨x₂, _⟩⟩, cc }) From 13a5329a8625701af92e9a96ffc90fa787fff24d Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Tue, 31 Jan 2023 13:11:30 +0000 Subject: [PATCH 0386/1291] chore(*): add mathlib4 synchronization comments (#18333) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.big_operators.finsupp` * `algebra.direct_sum.basic` * `algebra.is_prime_pow` * `combinatorics.set_family.shadow` * `data.dfinsupp.ne_locus` * `data.finsupp.fintype` * `data.finsupp.ne_locus` * `data.fun_like.fintype` * `group_theory.group_action.sub_mul_action.pointwise` * `group_theory.monoid_localization` * `group_theory.subgroup.finite` * `order.filter.modeq` * `order.filter.pointwise` * `ring_theory.subsemiring.basic` --- src/algebra/big_operators/finsupp.lean | 3 +++ src/algebra/direct_sum/basic.lean | 3 +++ src/algebra/is_prime_pow.lean | 3 +++ src/combinatorics/set_family/shadow.lean | 3 +++ src/data/dfinsupp/ne_locus.lean | 3 +++ src/data/finsupp/fintype.lean | 3 +++ src/data/finsupp/ne_locus.lean | 3 +++ src/data/fun_like/fintype.lean | 3 +++ src/group_theory/group_action/sub_mul_action/pointwise.lean | 3 +++ src/group_theory/monoid_localization.lean | 3 +++ src/group_theory/subgroup/finite.lean | 3 +++ src/order/filter/modeq.lean | 3 +++ src/order/filter/pointwise.lean | 3 +++ src/ring_theory/subsemiring/basic.lean | 3 +++ 14 files changed, 42 insertions(+) diff --git a/src/algebra/big_operators/finsupp.lean b/src/algebra/big_operators/finsupp.lean index ded7e78f17f9f..456f1fd9dcb51 100644 --- a/src/algebra/big_operators/finsupp.lean +++ b/src/algebra/big_operators/finsupp.lean @@ -13,6 +13,9 @@ import group_theory.submonoid.membership /-! # Big operators for finsupps +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains theorems relevant to big operators in finitely supported functions. -/ diff --git a/src/algebra/direct_sum/basic.lean b/src/algebra/direct_sum/basic.lean index 0b14312f4daea..181c4a6c7f016 100644 --- a/src/algebra/direct_sum/basic.lean +++ b/src/algebra/direct_sum/basic.lean @@ -8,6 +8,9 @@ import group_theory.submonoid.operations /-! # Direct sum +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the direct sum of abelian groups, indexed by a discrete type. ## Notation diff --git a/src/algebra/is_prime_pow.lean b/src/algebra/is_prime_pow.lean index 862dd75df9797..790f67bedab9b 100644 --- a/src/algebra/is_prime_pow.lean +++ b/src/algebra/is_prime_pow.lean @@ -9,6 +9,9 @@ import number_theory.divisors /-! # Prime powers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file deals with prime powers: numbers which are positive integer powers of a single prime. -/ diff --git a/src/combinatorics/set_family/shadow.lean b/src/combinatorics/set_family/shadow.lean index 7471273e79fb1..f02c1165296ea 100644 --- a/src/combinatorics/set_family/shadow.lean +++ b/src/combinatorics/set_family/shadow.lean @@ -9,6 +9,9 @@ import logic.function.iterate /-! # Shadows +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines shadows of a set family. The shadow of a set family is the set family of sets we get by removing any element from any set of the original family. If one pictures `finset α` as a big hypercube (each dimension being membership of a given element), then taking the shadow corresponds diff --git a/src/data/dfinsupp/ne_locus.lean b/src/data/dfinsupp/ne_locus.lean index 3d50a0f83d23a..772d04af470e8 100644 --- a/src/data/dfinsupp/ne_locus.lean +++ b/src/data/dfinsupp/ne_locus.lean @@ -8,6 +8,9 @@ import data.dfinsupp.basic /-! # Locus of unequal values of finitely supported dependent functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Let `N : α → Type*` be a type family, assume that `N a` has a `0` for all `a : α` and let `f g : Π₀ a, N a` be finitely supported dependent functions. diff --git a/src/data/finsupp/fintype.lean b/src/data/finsupp/fintype.lean index 0f77c4113b1dd..9f776641e39f0 100644 --- a/src/data/finsupp/fintype.lean +++ b/src/data/finsupp/fintype.lean @@ -10,6 +10,9 @@ import data.fintype.basic # Finiteness and infiniteness of `finsupp` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Some lemmas on the combination of `finsupp`, `fintype` and `infinite`. -/ diff --git a/src/data/finsupp/ne_locus.lean b/src/data/finsupp/ne_locus.lean index 39b84a7643a10..3a37840ccb78e 100644 --- a/src/data/finsupp/ne_locus.lean +++ b/src/data/finsupp/ne_locus.lean @@ -8,6 +8,9 @@ import data.finsupp.defs /-! # Locus of unequal values of finitely supported functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Let `α N` be two Types, assume that `N` has a `0` and let `f g : α →₀ N` be finitely supported functions. diff --git a/src/data/fun_like/fintype.lean b/src/data/fun_like/fintype.lean index 64f0b9e8732f8..168b503062b7b 100644 --- a/src/data/fun_like/fintype.lean +++ b/src/data/fun_like/fintype.lean @@ -11,6 +11,9 @@ import data.fun_like.basic /-! # Finiteness of `fun_like` types +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We show a type `F` with a `fun_like F α β` is finite if both `α` and `β` are finite. This corresponds to the following two pairs of declarations: diff --git a/src/group_theory/group_action/sub_mul_action/pointwise.lean b/src/group_theory/group_action/sub_mul_action/pointwise.lean index 3c233120b93d5..e54d5c31913a0 100644 --- a/src/group_theory/group_action/sub_mul_action/pointwise.lean +++ b/src/group_theory/group_action/sub_mul_action/pointwise.lean @@ -9,6 +9,9 @@ import group_theory.group_action.sub_mul_action /-! # Pointwise monoid structures on sub_mul_action +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file provides `sub_mul_action.monoid` and weaker typeclasses, which show that `sub_mul_action`s inherit the same pointwise multiplications as sets. diff --git a/src/group_theory/monoid_localization.lean b/src/group_theory/monoid_localization.lean index 2a89add0d748b..20c82192d58bb 100644 --- a/src/group_theory/monoid_localization.lean +++ b/src/group_theory/monoid_localization.lean @@ -10,6 +10,9 @@ import algebra.group.units /-! # Localizations of commutative monoids +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Localizing a commutative ring at one of its submonoids does not rely on the ring's addition, so we can generalize localizations to commutative monoids. diff --git a/src/group_theory/subgroup/finite.lean b/src/group_theory/subgroup/finite.lean index b8de0c882f0db..a5ab6c826172f 100644 --- a/src/group_theory/subgroup/finite.lean +++ b/src/group_theory/subgroup/finite.lean @@ -11,6 +11,9 @@ import group_theory.submonoid.membership /-! # Subgroups +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file provides some result on multiplicative and additive subgroups in the finite context. ## Tags diff --git a/src/order/filter/modeq.lean b/src/order/filter/modeq.lean index 94fee945148c6..015caefbdf3fe 100644 --- a/src/order/filter/modeq.lean +++ b/src/order/filter/modeq.lean @@ -9,6 +9,9 @@ import order.filter.at_top_bot /-! # Numbers are frequently modeq to fixed numbers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove that `m ≡ d [MOD n]` frequently as `m → ∞`. -/ diff --git a/src/order/filter/pointwise.lean b/src/order/filter/pointwise.lean index 4c4dca90e03fc..f9355b3a67674 100644 --- a/src/order/filter/pointwise.lean +++ b/src/order/filter/pointwise.lean @@ -10,6 +10,9 @@ import order.filter.ultrafilter /-! # Pointwise operations on filters +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines pointwise operations on filters. This is useful because usual algebraic operations distribute over pointwise operations. For example, * `(f₁ * f₂).map m = f₁.map m * f₂.map m` diff --git a/src/ring_theory/subsemiring/basic.lean b/src/ring_theory/subsemiring/basic.lean index 7fad17093710a..f06b9c51eff95 100644 --- a/src/ring_theory/subsemiring/basic.lean +++ b/src/ring_theory/subsemiring/basic.lean @@ -16,6 +16,9 @@ import group_theory.submonoid.membership /-! # Bundled subsemirings +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define bundled subsemirings and some standard constructions: `complete_lattice` structure, `subtype` and `inclusion` ring homomorphisms, subsemiring `map`, `comap` and range (`srange`) of a `ring_hom` etc. From b64b1f8825fd3d84743c1dc4256ae6f624dce612 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Tue, 31 Jan 2023 13:11:31 +0000 Subject: [PATCH 0387/1291] feat(algebra/periodic): generalize some lemmas (#18334) * Drop some commutativity assumptions. * Golf some proofs. * Add `protected` here and there for future compatibility with Lean 4. --- src/algebra/periodic.lean | 214 ++++++++++++++------------------------ 1 file changed, 77 insertions(+), 137 deletions(-) diff --git a/src/algebra/periodic.lean b/src/algebra/periodic.lean index 097404f758d10..c4a5570dd2f21 100644 --- a/src/algebra/periodic.lean +++ b/src/algebra/periodic.lean @@ -44,13 +44,11 @@ namespace function @[simp] def periodic [has_add α] (f : α → β) (c : α) : Prop := ∀ x : α, f (x + c) = f x -lemma periodic.funext [has_add α] - (h : periodic f c) : +protected lemma periodic.funext [has_add α] (h : periodic f c) : (λ x, f (x + c)) = f := funext h -lemma periodic.comp [has_add α] - (h : periodic f c) (g : β → γ) : +protected lemma periodic.comp [has_add α] (h : periodic f c) (g : β → γ) : periodic (g ∘ f) c := by simp * at * @@ -60,28 +58,24 @@ lemma periodic.comp_add_hom [has_add α] [has_add γ] λ x, by simp only [hg c, h (g x), add_hom.map_add, comp_app] @[to_additive] -lemma periodic.mul [has_add α] [has_mul β] - (hf : periodic f c) (hg : periodic g c) : +protected lemma periodic.mul [has_add α] [has_mul β] (hf : periodic f c) (hg : periodic g c) : periodic (f * g) c := by simp * at * @[to_additive] -lemma periodic.div [has_add α] [has_div β] - (hf : periodic f c) (hg : periodic g c) : +protected lemma periodic.div [has_add α] [has_div β] (hf : periodic f c) (hg : periodic g c) : periodic (f / g) c := by simp * at * @[to_additive] -lemma _root_.list.periodic_prod [has_add α] [comm_monoid β] +lemma _root_.list.periodic_prod [has_add α] [monoid β] (l : list (α → β)) (hl : ∀ f ∈ l, periodic f c) : periodic l.prod c := begin induction l with g l ih hl, { simp, }, - { simp only [list.mem_cons_iff, forall_eq_or_imp] at hl, - obtain ⟨hg, hl⟩ := hl, - simp only [list.prod_cons], - exact hg.mul (ih hl), }, + { rw [list.forall_mem_cons] at hl, + simpa only [list.prod_cons] using hl.1.mul (ih hl.2) } end @[to_additive] @@ -97,11 +91,11 @@ lemma _root_.finset.periodic_prod [has_add α] [comm_monoid β] s.prod_to_list f ▸ (s.to_list.map f).periodic_prod (by simpa [-periodic]) @[to_additive] -lemma periodic.smul [has_add α] [has_smul γ β] (h : periodic f c) (a : γ) : +protected lemma periodic.smul [has_add α] [has_smul γ β] (h : periodic f c) (a : γ) : periodic (a • f) c := by simp * at * -lemma periodic.const_smul [add_monoid α] [group γ] [distrib_mul_action γ α] +protected lemma periodic.const_smul [add_monoid α] [group γ] [distrib_mul_action γ α] (h : periodic f c) (a : γ) : periodic (λ x, f (a • x)) (a⁻¹ • c) := λ x, by simpa only [smul_add, smul_inv_smul] using h (a • x) @@ -115,8 +109,7 @@ begin simpa only [smul_add, smul_inv_smul₀ ha] using h (a • x), end -lemma periodic.const_mul [division_ring α] - (h : periodic f c) (a : α) : +protected lemma periodic.const_mul [division_ring α] (h : periodic f c) (a : α) : periodic (λ x, f (a * x)) (a⁻¹ * c) := h.const_smul₀ a @@ -130,13 +123,11 @@ lemma periodic.const_inv_smul₀ [add_comm_monoid α] [division_ring γ] [module periodic (λ x, f (a⁻¹ • x)) (a • c) := by simpa only [inv_inv] using h.const_smul₀ a⁻¹ -lemma periodic.const_inv_mul [division_ring α] - (h : periodic f c) (a : α) : +lemma periodic.const_inv_mul [division_ring α] (h : periodic f c) (a : α) : periodic (λ x, f (a⁻¹ * x)) (a * c) := h.const_inv_smul₀ a -lemma periodic.mul_const [division_ring α] - (h : periodic f c) (a : α) : +lemma periodic.mul_const [division_ring α] (h : periodic f c) (a : α) : periodic (λ x, f (x * a)) (c * a⁻¹) := h.const_smul₀ $ mul_opposite.op a @@ -145,40 +136,33 @@ lemma periodic.mul_const' [division_ring α] periodic (λ x, f (x * a)) (c / a) := by simpa only [div_eq_mul_inv] using h.mul_const a -lemma periodic.mul_const_inv [division_ring α] - (h : periodic f c) (a : α) : +lemma periodic.mul_const_inv [division_ring α] (h : periodic f c) (a : α) : periodic (λ x, f (x * a⁻¹)) (c * a) := h.const_inv_smul₀ $ mul_opposite.op a -lemma periodic.div_const [division_ring α] - (h : periodic f c) (a : α) : +lemma periodic.div_const [division_ring α] (h : periodic f c) (a : α) : periodic (λ x, f (x / a)) (c * a) := by simpa only [div_eq_mul_inv] using h.mul_const_inv a -lemma periodic.add_period [add_semigroup α] - (h1 : periodic f c₁) (h2 : periodic f c₂) : +lemma periodic.add_period [add_semigroup α] (h1 : periodic f c₁) (h2 : periodic f c₂) : periodic f (c₁ + c₂) := by simp [*, ← add_assoc] at * -lemma periodic.sub_eq [add_group α] - (h : periodic f c) (x : α) : +lemma periodic.sub_eq [add_group α] (h : periodic f c) (x : α) : f (x - c) = f x := by simpa only [sub_add_cancel] using (h (x - c)).symm -lemma periodic.sub_eq' [add_comm_group α] - (h : periodic f c) : +lemma periodic.sub_eq' [add_comm_group α] (h : periodic f c) : f (c - x) = f (-x) := by simpa only [sub_eq_neg_add] using h (-x) -lemma periodic.neg [add_group α] - (h : periodic f c) : +protected lemma periodic.neg [add_group α] (h : periodic f c) : periodic f (-c) := by simpa only [sub_eq_add_neg, periodic] using h.sub_eq -lemma periodic.sub_period [add_comm_group α] - (h1 : periodic f c₁) (h2 : periodic f c₂) : +lemma periodic.sub_period [add_group α] (h1 : periodic f c₁) (h2 : periodic f c₂) : periodic f (c₁ - c₂) := -let h := h2.neg in by simp [*, sub_eq_add_neg, add_comm c₁, ← add_assoc] at * +by simpa only [sub_eq_add_neg] using h1.add_period h2.neg lemma periodic.const_add [add_semigroup α] (h : periodic f c) (a : α) : periodic (λ x, f (a + x)) c := @@ -186,63 +170,49 @@ lemma periodic.const_add [add_semigroup α] (h : periodic f c) (a : α) : lemma periodic.add_const [add_comm_semigroup α] (h : periodic f c) (a : α) : periodic (λ x, f (x + a)) c := -λ x, by simpa [add_assoc x c a, add_comm c, ←add_assoc x a c] using h (x + a) +by simpa only [add_comm] using h.const_add a lemma periodic.const_sub [add_comm_group α] (h : periodic f c) (a : α) : periodic (λ x, f (a - x)) c := -begin - rw [←neg_neg c], - refine periodic.neg _, - intro x, - simpa [sub_add_eq_sub_sub] using h (a - x) -end +λ x, by simp only [← sub_sub, h.sub_eq] lemma periodic.sub_const [add_comm_group α] (h : periodic f c) (a : α) : periodic (λ x, f (x - a)) c := -λ x, by simpa [add_comm x c, add_sub_assoc, add_comm c (x - a)] using h (x - a) +by simpa only [sub_eq_add_neg] using h.add_const (-a) -lemma periodic.nsmul [add_monoid α] - (h : periodic f c) (n : ℕ) : +lemma periodic.nsmul [add_monoid α] (h : periodic f c) (n : ℕ) : periodic f (n • c) := by induction n; simp [nat.succ_eq_add_one, add_nsmul, ← add_assoc, zero_nsmul, *] at * -lemma periodic.nat_mul [semiring α] - (h : periodic f c) (n : ℕ) : +lemma periodic.nat_mul [semiring α] (h : periodic f c) (n : ℕ) : periodic f (n * c) := by simpa only [nsmul_eq_mul] using h.nsmul n -lemma periodic.neg_nsmul [add_group α] - (h : periodic f c) (n : ℕ) : +lemma periodic.neg_nsmul [add_group α] (h : periodic f c) (n : ℕ) : periodic f (-(n • c)) := (h.nsmul n).neg -lemma periodic.neg_nat_mul [ring α] - (h : periodic f c) (n : ℕ) : +lemma periodic.neg_nat_mul [ring α] (h : periodic f c) (n : ℕ) : periodic f (-(n * c)) := (h.nat_mul n).neg -lemma periodic.sub_nsmul_eq [add_group α] - (h : periodic f c) (n : ℕ) : +lemma periodic.sub_nsmul_eq [add_group α] (h : periodic f c) (n : ℕ) : f (x - n • c) = f x := by simpa only [sub_eq_add_neg] using h.neg_nsmul n x -lemma periodic.sub_nat_mul_eq [ring α] - (h : periodic f c) (n : ℕ) : +lemma periodic.sub_nat_mul_eq [ring α] (h : periodic f c) (n : ℕ) : f (x - n * c) = f x := by simpa only [nsmul_eq_mul] using h.sub_nsmul_eq n -lemma periodic.nsmul_sub_eq [add_comm_group α] - (h : periodic f c) (n : ℕ) : +lemma periodic.nsmul_sub_eq [add_comm_group α] (h : periodic f c) (n : ℕ) : f (n • c - x) = f (-x) := -by simpa only [sub_eq_neg_add] using h.nsmul n (-x) +(h.nsmul n).sub_eq' -lemma periodic.nat_mul_sub_eq [ring α] - (h : periodic f c) (n : ℕ) : +lemma periodic.nat_mul_sub_eq [ring α] (h : periodic f c) (n : ℕ) : f (n * c - x) = f (-x) := by simpa only [sub_eq_neg_add] using h.nat_mul n (-x) -lemma periodic.zsmul [add_group α] - (h : periodic f c) (n : ℤ) : +protected lemma periodic.zsmul [add_group α] (h : periodic f c) (n : ℤ) : periodic f (n • c) := begin cases n, @@ -250,58 +220,47 @@ begin { simpa only [zsmul_neg_succ_of_nat] using (h.nsmul n.succ).neg }, end -lemma periodic.int_mul [ring α] - (h : periodic f c) (n : ℤ) : +protected lemma periodic.int_mul [ring α] (h : periodic f c) (n : ℤ) : periodic f (n * c) := by simpa only [zsmul_eq_mul] using h.zsmul n -lemma periodic.sub_zsmul_eq [add_group α] - (h : periodic f c) (n : ℤ) : +lemma periodic.sub_zsmul_eq [add_group α] (h : periodic f c) (n : ℤ) : f (x - n • c) = f x := (h.zsmul n).sub_eq x -lemma periodic.sub_int_mul_eq [ring α] - (h : periodic f c) (n : ℤ) : +lemma periodic.sub_int_mul_eq [ring α] (h : periodic f c) (n : ℤ) : f (x - n * c) = f x := (h.int_mul n).sub_eq x -lemma periodic.zsmul_sub_eq [add_comm_group α] - (h : periodic f c) (n : ℤ) : +lemma periodic.zsmul_sub_eq [add_comm_group α] (h : periodic f c) (n : ℤ) : f (n • c - x) = f (-x) := -by simpa only [sub_eq_neg_add] using h.zsmul n (-x) +(h.zsmul _).sub_eq' -lemma periodic.int_mul_sub_eq [ring α] - (h : periodic f c) (n : ℤ) : +lemma periodic.int_mul_sub_eq [ring α] (h : periodic f c) (n : ℤ) : f (n * c - x) = f (-x) := -by simpa only [sub_eq_neg_add] using h.int_mul n (-x) +(h.int_mul _).sub_eq' -lemma periodic.eq [add_zero_class α] - (h : periodic f c) : +protected lemma periodic.eq [add_zero_class α] (h : periodic f c) : f c = f 0 := by simpa only [zero_add] using h 0 -lemma periodic.neg_eq [add_group α] - (h : periodic f c) : +protected lemma periodic.neg_eq [add_group α] (h : periodic f c) : f (-c) = f 0 := h.neg.eq -lemma periodic.nsmul_eq [add_monoid α] - (h : periodic f c) (n : ℕ) : +protected lemma periodic.nsmul_eq [add_monoid α] (h : periodic f c) (n : ℕ) : f (n • c) = f 0 := (h.nsmul n).eq -lemma periodic.nat_mul_eq [semiring α] - (h : periodic f c) (n : ℕ) : +lemma periodic.nat_mul_eq [semiring α] (h : periodic f c) (n : ℕ) : f (n * c) = f 0 := (h.nat_mul n).eq -lemma periodic.zsmul_eq [add_group α] - (h : periodic f c) (n : ℤ) : +lemma periodic.zsmul_eq [add_group α] (h : periodic f c) (n : ℤ) : f (n • c) = f 0 := (h.zsmul n).eq -lemma periodic.int_mul_eq [ring α] - (h : periodic f c) (n : ℤ) : +lemma periodic.int_mul_eq [ring α] (h : periodic f c) (n : ℤ) : f (n * c) = f 0 := (h.int_mul n).eq @@ -370,25 +329,22 @@ rfl @[simp] def antiperiodic [has_add α] [has_neg β] (f : α → β) (c : α) : Prop := ∀ x : α, f (x + c) = -f x -lemma antiperiodic.funext [has_add α] [has_neg β] - (h : antiperiodic f c) : +protected lemma antiperiodic.funext [has_add α] [has_neg β] (h : antiperiodic f c) : (λ x, f (x + c)) = -f := funext h -lemma antiperiodic.funext' [has_add α] [has_involutive_neg β] - (h : antiperiodic f c) : +protected lemma antiperiodic.funext' [has_add α] [has_involutive_neg β] (h : antiperiodic f c) : (λ x, -f (x + c)) = f := (eq_neg_iff_eq_neg.mp h.funext).symm /-- If a function is `antiperiodic` with antiperiod `c`, then it is also `periodic` with period `2 * c`. -/ -lemma antiperiodic.periodic [semiring α] [has_involutive_neg β] - (h : antiperiodic f c) : +protected lemma antiperiodic.periodic [semiring α] [has_involutive_neg β] (h : antiperiodic f c) : periodic f (2 * c) := by simp [two_mul, ← add_assoc, h _] -lemma antiperiodic.eq [add_zero_class α] [has_neg β] - (h : antiperiodic f c) : f c = -f 0 := +protected lemma antiperiodic.eq [add_zero_class α] [has_neg β] (h : antiperiodic f c) : + f c = -f 0 := by simpa only [zero_add] using h 0 lemma antiperiodic.nat_even_mul_periodic [semiring α] [has_involutive_neg β] @@ -411,36 +367,16 @@ lemma antiperiodic.int_odd_mul_antiperiodic [ring α] [has_involutive_neg β] antiperiodic f (n * (2 * c) + c) := λ x, by rw [← add_assoc, h, h.periodic.int_mul] -lemma antiperiodic.nat_mul_eq_of_eq_zero [comm_semiring α] [neg_zero_class β] - (h : antiperiodic f c) (hi : f 0 = 0) (n : ℕ) : - f (n * c) = 0 := -begin - induction n with k hk, - { simp [hi] }, - { simp [hk, add_mul, h (k * c)] } -end - -lemma antiperiodic.int_mul_eq_of_eq_zero [comm_ring α] [subtraction_monoid β] - (h : antiperiodic f c) (hi : f 0 = 0) (n : ℤ) : - f (n * c) = 0 := -begin - rcases int.even_or_odd n with ⟨k, rfl⟩ | ⟨k, rfl⟩; - have hk : (k : α) * (2 * c) = 2 * k * c := by rw [mul_left_comm, ← mul_assoc], - { simpa [← two_mul, hk, hi] using (h.int_even_mul_periodic k).eq }, - { simpa [add_mul, hk, hi] using (h.int_odd_mul_antiperiodic k).eq }, -end - lemma antiperiodic.sub_eq [add_group α] [has_involutive_neg β] (h : antiperiodic f c) (x : α) : f (x - c) = -f x := by simp only [eq_neg_iff_eq_neg.mp (h (x - c)), sub_add_cancel] -lemma antiperiodic.sub_eq' [add_comm_group α] [has_neg β] - (h : antiperiodic f c) : +lemma antiperiodic.sub_eq' [add_comm_group α] [has_neg β] (h : antiperiodic f c) : f (c - x) = -f (-x) := by simpa only [sub_eq_neg_add] using h (-x) -lemma antiperiodic.neg [add_group α] [has_involutive_neg β] +protected lemma antiperiodic.neg [add_group α] [has_involutive_neg β] (h : antiperiodic f c) : antiperiodic f (-c) := by simpa only [sub_eq_add_neg, antiperiodic] using h.sub_eq @@ -450,29 +386,33 @@ lemma antiperiodic.neg_eq [add_group α] [has_involutive_neg β] f (-c) = -f 0 := by simpa only [zero_add] using h.neg 0 +lemma antiperiodic.nat_mul_eq_of_eq_zero [ring α] [neg_zero_class β] + (h : antiperiodic f c) (hi : f 0 = 0) : ∀ n : ℕ, f (n * c) = 0 +| 0 := by rwa [nat.cast_zero, zero_mul] +| (n + 1) := by simp [add_mul, antiperiodic.nat_mul_eq_of_eq_zero n, h _] + +lemma antiperiodic.int_mul_eq_of_eq_zero [ring α] [subtraction_monoid β] + (h : antiperiodic f c) (hi : f 0 = 0) : ∀ n : ℤ, f (n * c) = 0 +| (n : ℕ) := by rwa [int.cast_coe_nat, h.nat_mul_eq_of_eq_zero] +| -[1+n] := by rw [int.cast_neg_succ_of_nat, neg_mul, ← mul_neg, h.neg.nat_mul_eq_of_eq_zero hi] + lemma antiperiodic.const_add [add_semigroup α] [has_neg β] (h : antiperiodic f c) (a : α) : antiperiodic (λ x, f (a + x)) c := λ x, by simpa [add_assoc] using h (a + x) lemma antiperiodic.add_const [add_comm_semigroup α] [has_neg β] (h : antiperiodic f c) (a : α) : antiperiodic (λ x, f (x + a)) c := -λ x, by simpa [add_assoc x c a, add_comm c, ←add_assoc x a c] using h (x + a) +λ x, by simpa only [add_right_comm] using h (x + a) lemma antiperiodic.const_sub [add_comm_group α] [has_involutive_neg β] (h : antiperiodic f c) - (a : α) : - antiperiodic (λ x, f (a - x)) c := -begin - rw [←neg_neg c], - refine antiperiodic.neg _, - intro x, - simpa [sub_add_eq_sub_sub] using h (a - x) -end + (a : α) : antiperiodic (λ x, f (a - x)) c := +λ x, by simp only [← sub_sub, h.sub_eq] lemma antiperiodic.sub_const [add_comm_group α] [has_neg β] (h : antiperiodic f c) (a : α) : antiperiodic (λ x, f (x - a)) c := -λ x, by simpa [add_comm x c, add_sub_assoc, add_comm c (x - a)] using h (x - a) +by simpa only [sub_eq_add_neg] using h.add_const (-a) -lemma antiperiodic.smul [has_add α] [monoid γ] [add_group β] [distrib_mul_action γ β] +protected lemma antiperiodic.smul [has_add α] [monoid γ] [add_group β] [distrib_mul_action γ β] (h : antiperiodic f c) (a : γ) : antiperiodic (a • f) c := by simp * at * @@ -522,47 +462,47 @@ lemma antiperiodic.mul_const_inv [division_ring α] [has_neg β] antiperiodic (λ x, f (x * a⁻¹)) (c * a) := h.const_inv_smul₀ $ (mul_opposite.op_ne_zero_iff a).mpr ha -lemma antiperiodic.div_inv [division_ring α] [has_neg β] +protected lemma antiperiodic.div_inv [division_ring α] [has_neg β] (h : antiperiodic f c) {a : α} (ha : a ≠ 0) : antiperiodic (λ x, f (x / a)) (c * a) := by simpa only [div_eq_mul_inv] using h.mul_const_inv ha -lemma antiperiodic.add [add_group α] [has_involutive_neg β] +protected lemma antiperiodic.add [add_group α] [has_involutive_neg β] (h1 : antiperiodic f c₁) (h2 : antiperiodic f c₂) : periodic f (c₁ + c₂) := by simp [*, ← add_assoc] at * -lemma antiperiodic.sub [add_comm_group α] [has_involutive_neg β] +protected lemma antiperiodic.sub [add_group α] [has_involutive_neg β] (h1 : antiperiodic f c₁) (h2 : antiperiodic f c₂) : periodic f (c₁ - c₂) := -let h := h2.neg in by simp [*, sub_eq_add_neg, add_comm c₁, ← add_assoc] at * +by simpa only [sub_eq_add_neg] using h1.add h2.neg lemma periodic.add_antiperiod [add_group α] [has_neg β] (h1 : periodic f c₁) (h2 : antiperiodic f c₂) : antiperiodic f (c₁ + c₂) := by simp [*, ← add_assoc] at * -lemma periodic.sub_antiperiod [add_comm_group α] [has_involutive_neg β] +lemma periodic.sub_antiperiod [add_group α] [has_involutive_neg β] (h1 : periodic f c₁) (h2 : antiperiodic f c₂) : antiperiodic f (c₁ - c₂) := -let h := h2.neg in by simp [*, sub_eq_add_neg, add_comm c₁, ← add_assoc] at * +by simpa only [sub_eq_add_neg] using h1.add_antiperiod h2.neg lemma periodic.add_antiperiod_eq [add_group α] [has_neg β] (h1 : periodic f c₁) (h2 : antiperiodic f c₂) : f (c₁ + c₂) = -f 0 := (h1.add_antiperiod h2).eq -lemma periodic.sub_antiperiod_eq [add_comm_group α] [has_involutive_neg β] +lemma periodic.sub_antiperiod_eq [add_group α] [has_involutive_neg β] (h1 : periodic f c₁) (h2 : antiperiodic f c₂) : f (c₁ - c₂) = -f 0 := (h1.sub_antiperiod h2).eq -lemma antiperiodic.mul [has_add α] [has_mul β] [has_distrib_neg β] +protected lemma antiperiodic.mul [has_add α] [has_mul β] [has_distrib_neg β] (hf : antiperiodic f c) (hg : antiperiodic g c) : periodic (f * g) c := by simp * at * -lemma antiperiodic.div [has_add α] [division_monoid β] [has_distrib_neg β] +protected lemma antiperiodic.div [has_add α] [division_monoid β] [has_distrib_neg β] (hf : antiperiodic f c) (hg : antiperiodic g c) : periodic (f / g) c := by simp [*, neg_div_neg_eq] at * From 78f647f8517f021d839a7553d5dc97e79b508dea Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Tue, 31 Jan 2023 16:05:37 +0000 Subject: [PATCH 0388/1291] chore(order/filter, topology/algebra): golf (#18097) --- src/order/filter/n_ary.lean | 37 ++++---------- src/topology/algebra/group/basic.lean | 72 +++++++++------------------ 2 files changed, 32 insertions(+), 77 deletions(-) diff --git a/src/order/filter/n_ary.lean b/src/order/filter/n_ary.lean index 9e67f2e016d1b..b4bca5c1a51e0 100644 --- a/src/order/filter/n_ary.lean +++ b/src/order/filter/n_ary.lean @@ -58,29 +58,15 @@ lemma map_prod_eq_map₂ (m : α → β → γ) (f : filter α) (g : filter β) filter.map (λ p : α × β, m p.1 p.2) (f ×ᶠ g) = map₂ m f g := begin ext s, - split, - { intro hmem, - rw filter.mem_map_iff_exists_image at hmem, - obtain ⟨s', hs', hsub⟩ := hmem, - rw filter.mem_prod_iff at hs', - obtain ⟨t, ht, t', ht', hsub'⟩ := hs', - refine ⟨t, t', ht, ht', _⟩, - rw ← set.image_prod, - exact subset_trans (set.image_subset (λ (p : α × β), m p.fst p.snd) hsub') hsub }, - { intro hmem, - rw mem_map₂_iff at hmem, - obtain ⟨t, t', ht, ht', hsub⟩ := hmem, - rw ← set.image_prod at hsub, - rw filter.mem_map_iff_exists_image, - exact ⟨t ×ˢ t', filter.prod_mem_prod ht ht', hsub⟩ }, + simp [mem_prod_iff, prod_subset_iff] end lemma map_prod_eq_map₂' (m : α × β → γ) (f : filter α) (g : filter β) : filter.map m (f ×ᶠ g) = map₂ (λ a b, m (a, b)) f g := -by { refine eq.trans _ (map_prod_eq_map₂ (curry m) f g), ext, simp } +(congr_arg2 _ (uncurry_curry m).symm rfl).trans (map_prod_eq_map₂ _ _ _) @[simp] lemma map₂_mk_eq_prod (f : filter α) (g : filter β) : map₂ prod.mk f g = f ×ᶠ g := -by ext; simp [mem_prod_iff] +by simp only [← map_prod_eq_map₂, prod.mk.eta, map_id'] -- lemma image2_mem_map₂_iff (hm : injective2 m) : image2 m s t ∈ map₂ m f g ↔ s ∈ f ∧ t ∈ g := -- ⟨by { rintro ⟨u, v, hu, hv, h⟩, rw image2_subset_image2_iff hm at h, @@ -230,19 +216,13 @@ begin end lemma map_map₂ (m : α → β → γ) (n : γ → δ) : (map₂ m f g).map n = map₂ (λ a b, n (m a b)) f g := -filter.ext $ λ u, exists₂_congr $ λ s t, by rw [←image_subset_iff, image_image2] +by rw [← map_prod_eq_map₂, ← map_prod_eq_map₂, map_map] lemma map₂_map_left (m : γ → β → δ) (n : α → γ) : map₂ m (f.map n) g = map₂ (λ a b, m (n a) b) f g := begin - ext u, - split, - { rintro ⟨s, t, hs, ht, hu⟩, - refine ⟨_, t, hs, ht, _⟩, - rw ←image2_image_left, - exact (image2_subset_right $ image_preimage_subset _ _).trans hu }, - { rintro ⟨s, t, hs, ht, hu⟩, - exact ⟨_, t, image_mem_map hs, ht, by rwa image2_image_left⟩ } + rw [← map_prod_eq_map₂, ← map_prod_eq_map₂, ← @map_id _ g, prod_map_map_eq, map_map, map_id], + refl end lemma map₂_map_right (m : α → γ → δ) (n : β → γ) : @@ -251,10 +231,11 @@ by rw [map₂_swap, map₂_map_left, map₂_swap] @[simp] lemma map₂_curry (m : α × β → γ) (f : filter α) (g : filter β) : map₂ (curry m) f g = (f ×ᶠ g).map m := -by { classical, rw [←map₂_mk_eq_prod, map_map₂, curry] } +(map_prod_eq_map₂' _ _ _).symm @[simp] lemma map_uncurry_prod (m : α → β → γ) (f : filter α) (g : filter β) : - (f ×ᶠ g).map (uncurry m) = map₂ m f g := by rw [←map₂_curry, curry_uncurry] + (f ×ᶠ g).map (uncurry m) = map₂ m f g := +by rw [←map₂_curry, curry_uncurry] /-! ### Algebraic replacement rules diff --git a/src/topology/algebra/group/basic.lean b/src/topology/algebra/group/basic.lean index 910414ccd7f62..c2b5c7ed806e5 100644 --- a/src/topology/algebra/group/basic.lean +++ b/src/topology/algebra/group/basic.lean @@ -670,24 +670,17 @@ lemma topological_group.ext_iff {G : Type*} [group G] {t t' : topological_space ⟨λ h, h ▸ rfl, tg.ext tg'⟩ @[to_additive] -lemma topological_group.of_nhds_aux {G : Type*} [group G] [topological_space G] +lemma has_continuous_inv.of_nhds_one {G : Type*} [group G] [topological_space G] (hinv : tendsto (λ (x : G), x⁻¹) (𝓝 1) (𝓝 1)) (hleft : ∀ (x₀ : G), 𝓝 x₀ = map (λ (x : G), x₀ * x) (𝓝 1)) - (hconj : ∀ (x₀ : G), map (λ (x : G), x₀ * x * x₀⁻¹) (𝓝 1) ≤ 𝓝 1) : continuous (λ x : G, x⁻¹) := + (hconj : ∀ (x₀ : G), tendsto (λ (x : G), x₀ * x * x₀⁻¹) (𝓝 1) (𝓝 1)) : + has_continuous_inv G := begin - rw continuous_iff_continuous_at, - rintros x₀, - have key : (λ x, (x₀*x)⁻¹) = (λ x, x₀⁻¹*x) ∘ (λ x, x₀*x*x₀⁻¹) ∘ (λ x, x⁻¹), - by {ext ; simp[mul_assoc] }, - calc map (λ x, x⁻¹) (𝓝 x₀) - = map (λ x, x⁻¹) (map (λ x, x₀*x) $ 𝓝 1) : by rw hleft - ... = map (λ x, (x₀*x)⁻¹) (𝓝 1) : by rw filter.map_map - ... = map (((λ x, x₀⁻¹*x) ∘ (λ x, x₀*x*x₀⁻¹)) ∘ (λ x, x⁻¹)) (𝓝 1) : by rw key - ... = map ((λ x, x₀⁻¹*x) ∘ (λ x, x₀*x*x₀⁻¹)) _ : by rw ← filter.map_map - ... ≤ map ((λ x, x₀⁻¹ * x) ∘ λ x, x₀ * x * x₀⁻¹) (𝓝 1) : map_mono hinv - ... = map (λ x, x₀⁻¹ * x) (map (λ x, x₀ * x * x₀⁻¹) (𝓝 1)) : filter.map_map - ... ≤ map (λ x, x₀⁻¹ * x) (𝓝 1) : map_mono (hconj x₀) - ... = 𝓝 x₀⁻¹ : (hleft _).symm + refine ⟨continuous_iff_continuous_at.2 $ λ x₀, _⟩, + have : tendsto (λ x, x₀⁻¹ * (x₀ * x⁻¹ * x₀⁻¹)) (𝓝 1) (map ((*) x₀⁻¹) (𝓝 1)), + from (tendsto_map.comp $ hconj x₀).comp hinv, + simpa only [continuous_at, hleft x₀, hleft x₀⁻¹, tendsto_map'_iff, (∘), mul_assoc, + mul_inv_rev, inv_mul_cancel_left] using this end @[to_additive] @@ -696,17 +689,13 @@ lemma topological_group.of_nhds_one' {G : Type u} [group G] [topological_space G (hinv : tendsto (λ x : G, x⁻¹) (𝓝 1) (𝓝 1)) (hleft : ∀ x₀ : G, 𝓝 x₀ = map (λ x, x₀*x) (𝓝 1)) (hright : ∀ x₀ : G, 𝓝 x₀ = map (λ x, x*x₀) (𝓝 1)) : topological_group G := -begin - refine { continuous_mul := (has_continuous_mul.of_nhds_one hmul hleft hright).continuous_mul, - continuous_inv := topological_group.of_nhds_aux hinv hleft _ }, - intros x₀, - suffices : map (λ (x : G), x₀ * x * x₀⁻¹) (𝓝 1) = 𝓝 1, by simp [this, le_refl], - rw [show (λ x, x₀ * x * x₀⁻¹) = (λ x, x₀ * x) ∘ λ x, x*x₀⁻¹, by {ext, simp [mul_assoc] }, - ← filter.map_map, ← hright, hleft x₀⁻¹, filter.map_map], - convert map_id, - ext, - simp -end +{ to_has_continuous_mul := has_continuous_mul.of_nhds_one hmul hleft hright, + to_has_continuous_inv := has_continuous_inv.of_nhds_one hinv hleft $ λ x₀, le_of_eq + begin + rw [show (λ x, x₀ * x * x₀⁻¹) = (λ x, x * x₀⁻¹) ∘ (λ x, x₀ * x), from rfl, ← map_map, + ← hleft, hright, map_map], + simp [(∘)] + end } @[to_additive] lemma topological_group.of_nhds_one {G : Type u} [group G] [topological_space G] @@ -714,29 +703,14 @@ lemma topological_group.of_nhds_one {G : Type u} [group G] [topological_space G] (hinv : tendsto (λ x : G, x⁻¹) (𝓝 1) (𝓝 1)) (hleft : ∀ x₀ : G, 𝓝 x₀ = map (λ x, x₀*x) (𝓝 1)) (hconj : ∀ x₀ : G, tendsto (λ x, x₀*x*x₀⁻¹) (𝓝 1) (𝓝 1)) : topological_group G := - { continuous_mul := begin - rw continuous_iff_continuous_at, - rintros ⟨x₀, y₀⟩, - have key : (λ (p : G × G), x₀ * p.1 * (y₀ * p.2)) = - ((λ x, x₀*y₀*x) ∘ (uncurry (*)) ∘ (prod.map (λ x, y₀⁻¹*x*y₀) id)), - by { ext, simp [uncurry, prod.map, mul_assoc] }, - specialize hconj y₀⁻¹, rw inv_inv at hconj, - calc map (λ (p : G × G), p.1 * p.2) (𝓝 (x₀, y₀)) - = map (λ (p : G × G), p.1 * p.2) ((𝓝 x₀) ×ᶠ 𝓝 y₀) - : by rw nhds_prod_eq - ... = map (λ (p : G × G), x₀ * p.1 * (y₀ * p.2)) ((𝓝 1) ×ᶠ (𝓝 1)) - : by rw [hleft x₀, hleft y₀, prod_map_map_eq, filter.map_map] - ... = map (((λ x, x₀*y₀*x) ∘ (uncurry (*))) ∘ (prod.map (λ x, y₀⁻¹*x*y₀) id))((𝓝 1) ×ᶠ (𝓝 1)) - : by rw key - ... = map ((λ x, x₀*y₀*x) ∘ (uncurry (*))) ((map (λ x, y₀⁻¹*x*y₀) $ 𝓝 1) ×ᶠ (𝓝 1)) - : by rw [← filter.map_map, ← prod_map_map_eq', map_id] - ... ≤ map ((λ x, x₀*y₀*x) ∘ (uncurry (*))) ((𝓝 1) ×ᶠ (𝓝 1)) - : map_mono (filter.prod_mono hconj $ le_rfl) - ... = map (λ x, x₀*y₀*x) (map (uncurry (*)) ((𝓝 1) ×ᶠ (𝓝 1))) : by rw filter.map_map - ... ≤ map (λ x, x₀*y₀*x) (𝓝 1) : map_mono hmul - ... = 𝓝 (x₀*y₀) : (hleft _).symm - end, - continuous_inv := topological_group.of_nhds_aux hinv hleft hconj} +begin + refine topological_group.of_nhds_one' hmul hinv hleft (λ x₀, _), + replace hconj : ∀ x₀ : G, map (λ x, x₀ * x * x₀⁻¹) (𝓝 1) = 𝓝 1, + from λ x₀, map_eq_of_inverse (λ x, x₀⁻¹ * x * x₀⁻¹⁻¹) (by { ext, simp [mul_assoc] }) + (hconj _) (hconj _), + rw [← hconj x₀], + simpa [(∘)] using hleft _ +end @[to_additive] lemma topological_group.of_comm_of_nhds_one {G : Type u} [comm_group G] [topological_space G] From 8233a1cb983005c6863fcf6de53408871eec06fe Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Tue, 31 Jan 2023 16:05:38 +0000 Subject: [PATCH 0389/1291] feat(ring_theory): definition and properties of relative ideal norm (#18299) This PR provides a definition of the relative ideal norm `ideal.span_norm R I` as the ideal spanned by the norms of elements in `I`. We also give some basic results on this definition. A follow-up PR will bundle this into a homomorphism `ideal.rel_norm`. --- src/ring_theory/ideal/norm.lean | 98 +++++++++++++++++++++++++++++++-- 1 file changed, 92 insertions(+), 6 deletions(-) diff --git a/src/ring_theory/ideal/norm.lean b/src/ring_theory/ideal/norm.lean index 8cb9be9f0ad18..deb0ebf48daf5 100644 --- a/src/ring_theory/ideal/norm.lean +++ b/src/ring_theory/ideal/norm.lean @@ -14,14 +14,16 @@ import linear_algebra.free_module.ideal_quotient import linear_algebra.free_module.pid import linear_algebra.isomorphisms import ring_theory.dedekind_domain.ideal -import ring_theory.norm +import ring_theory.localization.norm /-! # Ideal norms This file defines the absolute ideal norm `ideal.abs_norm (I : ideal R) : ℕ` as the cardinality of -the quotient `R ⧸ I` (setting it to 0 if the cardinality is infinite). +the quotient `R ⧸ I` (setting it to 0 if the cardinality is infinite), +and the relative ideal norm `ideal.span_norm R (I : ideal S) : ideal S` as the ideal spanned by +the norms of elements in `I`. ## Main definitions @@ -29,6 +31,8 @@ the quotient `R ⧸ I` (setting it to 0 if the cardinality is infinite). This maps `⊥` to `0` and `⊤` to `1`. * `ideal.abs_norm (I : ideal R)`: the absolute ideal norm, defined as the cardinality of the quotient `R ⧸ I`, as a bundled monoid-with-zero homomorphism. + * `ideal.span_norm R (I : ideal S)`: the relative ideal norm, defined as + the ideal spanned by the norms of elements in `I`. ## Main results @@ -38,13 +42,12 @@ the quotient `R ⧸ I` (setting it to 0 if the cardinality is infinite). of the basis change matrix * `ideal.abs_norm_span_singleton`: the ideal norm of a principal ideal is the norm of its generator - -## TODO - -Define the relative norm. -/ open_locale big_operators +open_locale non_zero_divisors + +section abs_norm namespace submodule @@ -408,3 +411,86 @@ lemma prime_of_irreducible_abs_norm_span {a : S} (ha : a ≠ 0) (ideal.span_singleton_prime ha).mp (is_prime_of_irreducible_abs_norm hI) end ideal + +end abs_norm + +section span_norm + +namespace ideal + +open submodule + +variables (R : Type*) [comm_ring R] {S : Type*} [comm_ring S] [algebra R S] + +/-- `ideal.span_norm R (I : ideal S)` is the ideal generated by mapping `algebra.norm R` over `I`. +-/ +def span_norm (I : ideal S) : ideal R := +ideal.span (algebra.norm R '' (I : set S)) + +@[simp] lemma span_norm_bot + [nontrivial S] [module.free R S] [module.finite R S] : + span_norm R (⊥ : ideal S) = ⊥ := +span_eq_bot.mpr (λ x hx, by simpa using hx) + +lemma norm_mem_span_norm {I : ideal S} (x : S) (hx : x ∈ I) : algebra.norm R x ∈ I.span_norm R := +subset_span (set.mem_image_of_mem _ hx) + +@[simp] lemma span_norm_singleton {r : S} : + span_norm R (span ({r} : set S)) = span {algebra.norm R r} := +le_antisymm + (span_le.mpr (λ x hx, mem_span_singleton.mpr begin + obtain ⟨x, hx', rfl⟩ := (set.mem_image _ _ _).mp hx, + exact map_dvd _ (mem_span_singleton.mp hx') + end)) + ((span_singleton_le_iff_mem _).mpr (norm_mem_span_norm _ _ (mem_span_singleton_self _))) + +@[simp] lemma span_norm_top : span_norm R (⊤ : ideal S) = 1 := +by simp [← ideal.span_singleton_one] + +lemma map_span_norm (I : ideal S) {T : Type*} [comm_ring T] (f : R →+* T) : + map f (span_norm R I) = span ((f ∘ algebra.norm R) '' (I : set S)) := +by rw [span_norm, map_span, set.image_image] + +@[mono] +lemma span_norm_mono {I J : ideal S} (h : I ≤ J) : span_norm R I ≤ span_norm R J := +ideal.span_mono (set.monotone_image h) + +lemma span_norm_localization (I : ideal S) [module.finite R S] [module.free R S] + {M : submonoid R} {Rₘ Sₘ : Type*} + [comm_ring Rₘ] [algebra R Rₘ] [comm_ring Sₘ] [algebra S Sₘ] + [algebra Rₘ Sₘ] [algebra R Sₘ] [is_scalar_tower R Rₘ Sₘ] [is_scalar_tower R S Sₘ] + [is_localization M Rₘ] [is_localization (algebra.algebra_map_submonoid S M) Sₘ] + (hM : algebra.algebra_map_submonoid S M ≤ S⁰) : + span_norm Rₘ (I.map (algebra_map S Sₘ)) = (span_norm R I).map (algebra_map R Rₘ) := +begin + casesI h : subsingleton_or_nontrivial R, + { haveI := is_localization.unique R Rₘ M, + simp }, + let b := module.free.choose_basis R S, + rw map_span_norm, + refine span_eq_span (set.image_subset_iff.mpr _) (set.image_subset_iff.mpr _), + { rintros a' ha', + simp only [set.mem_preimage, submodule_span_eq, ← map_span_norm, set_like.mem_coe, + is_localization.mem_map_algebra_map_iff (algebra.algebra_map_submonoid S M) Sₘ, + is_localization.mem_map_algebra_map_iff M Rₘ, prod.exists] + at ⊢ ha', + obtain ⟨⟨a, ha⟩, ⟨_, ⟨s, hs, rfl⟩⟩, has⟩ := ha', + refine ⟨⟨algebra.norm R a, norm_mem_span_norm _ _ ha⟩, + ⟨s ^ fintype.card (module.free.choose_basis_index R S), pow_mem hs _⟩, _⟩, + swap, + simp only [submodule.coe_mk, subtype.coe_mk, map_pow] at ⊢ has, + apply_fun algebra.norm Rₘ at has, + rwa [_root_.map_mul, ← is_scalar_tower.algebra_map_apply, + is_scalar_tower.algebra_map_apply R Rₘ, + algebra.norm_algebra_map_of_basis (b.localization_localization Rₘ M Sₘ hM), + algebra.norm_localization R a hM] at has, + all_goals { apply_instance} }, + { intros a ha, + rw [set.mem_preimage, function.comp_app, ← algebra.norm_localization R a hM], + exact subset_span (set.mem_image_of_mem _ (mem_map_of_mem _ ha)), + all_goals { apply_instance} }, +end + +end ideal + +end span_norm From afdb43429311b885a7988ea15d0bac2aac80f69c Mon Sep 17 00:00:00 2001 From: Michael Stoll Date: Tue, 31 Jan 2023 16:05:39 +0000 Subject: [PATCH 0390/1291] =?UTF-8?q?feat(algebra/order/floor):=20`0=20 Date: Tue, 31 Jan 2023 16:05:41 +0000 Subject: [PATCH 0391/1291] feat(data/quot): make `trunc` a `quotient` (#18320) This allows us to use `quotient` lemmas for `trunc`. mathlib4 PR: https://github.com/leanprover-community/mathlib4/pull/1924 Co-authored-by: Eric Wieser --- src/data/quot.lean | 14 ++++++++++---- 1 file changed, 10 insertions(+), 4 deletions(-) diff --git a/src/data/quot.lean b/src/data/quot.lean index eb8bc111569c6..fac3f392436ed 100644 --- a/src/data/quot.lean +++ b/src/data/quot.lean @@ -365,15 +365,21 @@ lemma nonempty_quotient_iff (s : setoid α) : nonempty (quotient s) ↔ nonempty /-! ### Truncation -/ +theorem true_equivalence : @equivalence α (λ _ _, true) := +⟨λ _, trivial, λ _ _ _, trivial, λ _ _ _ _ _, trivial⟩ + +/-- Always-true relation as a `setoid`. + +Note that in later files the preferred spelling is `⊤ : setoid α`. -/ +def true_setoid : setoid α := +⟨_, true_equivalence⟩ + /-- `trunc α` is the quotient of `α` by the always-true relation. This is related to the propositional truncation in HoTT, and is similar in effect to `nonempty α`, but unlike `nonempty α`, `trunc α` is data, so the VM representation is the same as `α`, and so this can be used to maintain computability. -/ -def {u} trunc (α : Sort u) : Sort u := @quot α (λ _ _, true) - -theorem true_equivalence : @equivalence α (λ _ _, true) := -⟨λ _, trivial, λ _ _ _, trivial, λ _ _ _ _ _, trivial⟩ +def {u} trunc (α : Sort u) : Sort u := @quotient α true_setoid namespace trunc From 2769416a2ebdfb072f0e0907bac7e6e7117dcf8c Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Tue, 31 Jan 2023 16:05:42 +0000 Subject: [PATCH 0392/1291] chore(order/upper_lower): split file (#18330) This way we can port some `topology.*` files without wating for leanprover-community/mathlib4#1636 Also delete some lemmas that are true for any `order_iso`. --- src/algebra/order/upper_lower.lean | 2 +- .../set_family/harris_kleitman.lean | 2 +- .../set_family/intersecting.lean | 2 +- src/combinatorics/young/young_diagram.lean | 2 +- src/order/countable_dense_linear_order.lean | 1 + src/order/ideal.lean | 2 +- src/order/minimal.lean | 2 +- .../basic.lean} | 70 +------------------ src/order/upper_lower/hom.lean | 65 +++++++++++++++++ src/topology/order/lower_topology.lean | 1 + src/topology/order/priestley.lean | 2 +- src/topology/sets/order.lean | 2 +- 12 files changed, 78 insertions(+), 75 deletions(-) rename src/order/{upper_lower.lean => upper_lower/basic.lean} (92%) create mode 100644 src/order/upper_lower/hom.lean diff --git a/src/algebra/order/upper_lower.lean b/src/algebra/order/upper_lower.lean index 9cb3be7ba60e7..7bfeea35e4f51 100644 --- a/src/algebra/order/upper_lower.lean +++ b/src/algebra/order/upper_lower.lean @@ -5,7 +5,7 @@ Authors: Yaël Dillies -/ import algebra.order.group.defs import data.set.pointwise.smul -import order.upper_lower +import order.upper_lower.basic /-! # Algebraic operations on upper/lower sets diff --git a/src/combinatorics/set_family/harris_kleitman.lean b/src/combinatorics/set_family/harris_kleitman.lean index bfd2055a2aee4..c4654dd7dc7a0 100644 --- a/src/combinatorics/set_family/harris_kleitman.lean +++ b/src/combinatorics/set_family/harris_kleitman.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ import combinatorics.set_family.compression.down -import order.upper_lower +import order.upper_lower.basic import data.fintype.big_operators /-! diff --git a/src/combinatorics/set_family/intersecting.lean b/src/combinatorics/set_family/intersecting.lean index ce2503b935cd1..f90a3cb4e0781 100644 --- a/src/combinatorics/set_family/intersecting.lean +++ b/src/combinatorics/set_family/intersecting.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ import data.fintype.card -import order.upper_lower +import order.upper_lower.basic /-! # Intersecting families diff --git a/src/combinatorics/young/young_diagram.lean b/src/combinatorics/young/young_diagram.lean index 5cc19d6b5eace..ed1174f80822d 100644 --- a/src/combinatorics/young/young_diagram.lean +++ b/src/combinatorics/young/young_diagram.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Jake Levinson. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jake Levinson -/ -import order.upper_lower +import order.upper_lower.basic import data.finset.preimage /-! diff --git a/src/order/countable_dense_linear_order.lean b/src/order/countable_dense_linear_order.lean index e2e803c076899..36ec280490d45 100644 --- a/src/order/countable_dense_linear_order.lean +++ b/src/order/countable_dense_linear_order.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: David Wärn -/ import order.ideal +import data.finset.lattice /-! # The back and forth method and countable dense linear orders diff --git a/src/order/ideal.lean b/src/order/ideal.lean index 462aa85b63c45..f979b12efec09 100644 --- a/src/order/ideal.lean +++ b/src/order/ideal.lean @@ -5,7 +5,7 @@ Authors: David Wärn -/ import logic.encodable.basic import order.atoms -import order.upper_lower +import order.upper_lower.basic /-! # Order ideals, cofinal sets, and the Rasiowa–Sikorski lemma diff --git a/src/order/minimal.lean b/src/order/minimal.lean index 2a84be94989e8..f898c52d85c3e 100644 --- a/src/order/minimal.lean +++ b/src/order/minimal.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ import order.antichain -import order.upper_lower +import order.upper_lower.basic /-! # Minimal/maximal elements of a set diff --git a/src/order/upper_lower.lean b/src/order/upper_lower/basic.lean similarity index 92% rename from src/order/upper_lower.lean rename to src/order/upper_lower/basic.lean index 2bcc1e90c8a6b..262f5b7c9283e 100644 --- a/src/order/upper_lower.lean +++ b/src/order/upper_lower/basic.lean @@ -6,7 +6,6 @@ Authors: Yaël Dillies, Sara Rousta import data.set_like.basic import data.set.intervals.ord_connected import data.set.intervals.order_iso -import order.hom.complete_lattice /-! # Up-sets and down-sets @@ -548,22 +547,6 @@ variables (f s t) @[simp, norm_cast] lemma coe_map : (map f s : set β) = f '' s := rfl -@[simp] protected lemma map_sup : map f (s ⊔ t) = map f s ⊔ map f t := ext $ image_inter f.injective -@[simp] protected lemma map_inf : map f (s ⊓ t) = map f s ⊓ map f t := ext $ image_union _ _ _ -@[simp] protected lemma map_top : map f ⊤ = ⊤ := ext $ image_empty _ -@[simp] protected lemma map_bot : map f ⊥ = ⊥ := ext $ image_univ_of_surjective f.surjective -@[simp] protected lemma map_Sup (S : set (upper_set α)) : map f (Sup S) = ⨆ s ∈ S, map f s := -ext $ by { push_cast, exact image_Inter₂ f.bijective _ } - -@[simp] protected lemma map_Inf (S : set (upper_set α)) : map f (Inf S) = ⨅ s ∈ S, map f s := -ext $ by { push_cast, exact image_Union₂ _ _ } - -@[simp] protected lemma map_supr (g : ι → upper_set α) : map f (⨆ i, g i) = ⨆ i, map f (g i) := -ext $ by { push_cast, exact image_Inter f.bijective _ } - -@[simp] protected lemma map_infi (g : ι → upper_set α) : map f (⨅ i, g i) = ⨅ i, map f (g i) := -ext $ by { push_cast, exact image_Union } - end upper_set namespace lower_set @@ -592,22 +575,6 @@ variables (f s t) @[simp, norm_cast] lemma coe_map : (map f s : set β) = f '' s := rfl -@[simp] protected lemma map_sup : map f (s ⊔ t) = map f s ⊔ map f t := ext $ image_union _ _ _ -@[simp] protected lemma map_inf : map f (s ⊓ t) = map f s ⊓ map f t := ext $ image_inter f.injective -@[simp] protected lemma map_top : map f ⊤ = ⊤ := ext $ image_univ_of_surjective f.surjective -@[simp] protected lemma map_bot : map f ⊥ = ⊥ := ext $ image_empty _ -@[simp] protected lemma map_Sup (S : set (lower_set α)) : map f (Sup S) = ⨆ s ∈ S, map f s := -ext $ by { push_cast, exact image_Union₂ _ _ } - -protected lemma map_Inf (S : set (lower_set α)) : map f (Inf S) = ⨅ s ∈ S, map f s := -ext $ by { push_cast, exact image_Inter₂ f.bijective _ } - -protected lemma map_supr (g : ι → lower_set α) : map f (⨆ i, g i) = ⨆ i, map f (g i) := -ext $ by { push_cast, exact image_Union } - -protected lemma map_infi (g : ι → lower_set α) : map f (⨅ i, g i) = ⨅ i, map f (g i) := -ext $ by { push_cast, exact image_Inter f.bijective _ } - end lower_set namespace upper_set @@ -654,17 +621,8 @@ lemma Ici_le_Ioi (a : α) : Ici a ≤ Ioi a := Ioi_subset_Ici_self end preorder -section semilattice_sup -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 α (upper_set α) := ⟨Ici, Ici_sup⟩ - -@[simp] lemma Ici_sup_hom_apply (a : α) : Ici_sup_hom a = (Ici a) := rfl - -end semilattice_sup +@[simp] lemma Ici_sup [semilattice_sup α] (a b : α) : Ici (a ⊔ b) = Ici a ⊔ Ici b := +ext Ici_inter_Ici.symm section complete_lattice variables [complete_lattice α] @@ -678,11 +636,6 @@ set_like.ext $ λ c, by simp only [mem_Ici_iff, mem_supr_iff, supr_le_iff] @[simp] lemma Ici_supr₂ (f : Π i, κ i → α) : Ici (⨆ i j, f i j) = ⨆ i j, Ici (f i j) := by simp_rw Ici_supr -/-- `upper_set.Ici` as a `Sup_hom`. -/ -def Ici_Sup_hom : Sup_hom α (upper_set α) := ⟨Ici, λ s, (Ici_Sup s).trans Sup_image.symm⟩ - -@[simp] lemma Ici_Sup_hom_apply (a : α) : Ici_Sup_hom a = to_dual (Ici a) := rfl - end complete_lattice end upper_set @@ -711,20 +664,9 @@ lemma Ioi_le_Ici (a : α) : Ioi a ≤ Ici a := Ioi_subset_Ici_self end preorder -section semilattice_inf -variables [semilattice_inf α] - -@[simp] lemma Iic_inf (a b : α) : Iic (a ⊓ b) = Iic a ⊓ Iic b := +@[simp] lemma Iic_inf [semilattice_inf α] (a b : α) : Iic (a ⊓ b) = Iic a ⊓ Iic b := set_like.coe_injective Iic_inter_Iic.symm -/-- `lower_set.Iic` as an `inf_hom`. -/ -def Iic_inf_hom : inf_hom α (lower_set α) := ⟨Iic, Iic_inf⟩ - -@[simp] lemma coe_Iic_inf_hom : (Iic_inf_hom : α → lower_set α) = Iic := rfl -@[simp] lemma Iic_inf_hom_apply (a : α) : Iic_inf_hom a = Iic a := rfl - -end semilattice_inf - section complete_lattice variables [complete_lattice α] @@ -737,12 +679,6 @@ set_like.ext $ λ c, by simp only [mem_Iic_iff, mem_infi_iff, le_infi_iff] @[simp] lemma Iic_infi₂ (f : Π i, κ i → α) : Iic (⨅ i j, f i j) = ⨅ i j, Iic (f i j) := by simp_rw Iic_infi -/-- `lower_set.Iic` as an `Inf_hom`. -/ -def Iic_Inf_hom : Inf_hom α (lower_set α) := ⟨Iic, λ s, (Iic_Inf s).trans Inf_image.symm⟩ - -@[simp] lemma coe_Iic_Inf_hom : (Iic_Inf_hom : α → lower_set α) = Iic := rfl -@[simp] lemma Iic_Inf_hom_apply (a : α) : Iic_Inf_hom a = Iic a := rfl - end complete_lattice end lower_set diff --git a/src/order/upper_lower/hom.lean b/src/order/upper_lower/hom.lean new file mode 100644 index 0000000000000..285f0f407b18c --- /dev/null +++ b/src/order/upper_lower/hom.lean @@ -0,0 +1,65 @@ +/- +Copyright (c) 2022 Yaël Dillies, Sara Rousta. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies +-/ +import order.upper_lower.basic +import order.hom.complete_lattice + +/-! +# `upper_set.Ici` etc as `sup`/`Sup`/`inf`/`Inf`-homomorphisms + +In this file we define `upper_set.Ici_sup_hom` etc. These functions are `upper_set.Ici` and +`lower_set.Iic` bundled as `sup_hom`s, `inf_hom`s, `Sup_hom`s, or `Inf_hom`s. +-/ + +variable {α : Type*} +open order_dual + +namespace upper_set + +section semilattice_sup + +variable [semilattice_sup α] + +/-- `upper_set.Ici` as a `sup_hom`. -/ +def Ici_sup_hom : sup_hom α (upper_set α) := ⟨Ici, Ici_sup⟩ + +@[simp] lemma coe_Ici_sup_hom : (Ici_sup_hom : α → upper_set α) = Ici := rfl +@[simp] lemma Ici_sup_hom_apply (a : α) : Ici_sup_hom a = (Ici a) := rfl + +end semilattice_sup + +variable [complete_lattice α] + +/-- `upper_set.Ici` as a `Sup_hom`. -/ +def Ici_Sup_hom : Sup_hom α (upper_set α) := ⟨Ici, λ s, (Ici_Sup s).trans Sup_image.symm⟩ + +@[simp] lemma coe_Ici_Sup_hom : (Ici_Sup_hom : α → upper_set α) = Ici := rfl +@[simp] lemma Ici_Sup_hom_apply (a : α) : Ici_Sup_hom a = Ici a := rfl + +end upper_set + +namespace lower_set + +section semilattice_inf + +variable [semilattice_inf α] + +/-- `lower_set.Iic` as an `inf_hom`. -/ +def Iic_inf_hom : inf_hom α (lower_set α) := ⟨Iic, Iic_inf⟩ + +@[simp] lemma coe_Iic_inf_hom : (Iic_inf_hom : α → lower_set α) = Iic := rfl +@[simp] lemma Iic_inf_hom_apply (a : α) : Iic_inf_hom a = Iic a := rfl + +end semilattice_inf + +variable [complete_lattice α] + +/-- `lower_set.Iic` as an `Inf_hom`. -/ +def Iic_Inf_hom : Inf_hom α (lower_set α) := ⟨Iic, λ s, (Iic_Inf s).trans Inf_image.symm⟩ + +@[simp] lemma coe_Iic_Inf_hom : (Iic_Inf_hom : α → lower_set α) = Iic := rfl +@[simp] lemma Iic_Inf_hom_apply (a : α) : Iic_Inf_hom a = Iic a := rfl + +end lower_set diff --git a/src/topology/order/lower_topology.lean b/src/topology/order/lower_topology.lean index 0412ba2aa4092..496c12b200bb2 100644 --- a/src/topology/order/lower_topology.lean +++ b/src/topology/order/lower_topology.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Christopher Hoskin -/ import topology.order.lattice +import order.hom.complete_lattice /-! # Lower topology diff --git a/src/topology/order/priestley.lean b/src/topology/order/priestley.lean index 6a4688db360a1..15355ae5006c2 100644 --- a/src/topology/order/priestley.lean +++ b/src/topology/order/priestley.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ -import order.upper_lower +import order.upper_lower.basic import topology.separation /-! diff --git a/src/topology/sets/order.lean b/src/topology/sets/order.lean index a1a05edf4a398..d44b1dc3f74e4 100644 --- a/src/topology/sets/order.lean +++ b/src/topology/sets/order.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ -import order.upper_lower +import order.upper_lower.basic import topology.sets.closeds /-! From 24e75f1ee89ff37e99581084704f3f6a950db2ea Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Tue, 31 Jan 2023 16:05:43 +0000 Subject: [PATCH 0393/1291] chore(order/filter/basic): protect 3 lemmas (#18331) --- src/order/filter/basic.lean | 11 ++++++----- 1 file changed, 6 insertions(+), 5 deletions(-) diff --git a/src/order/filter/basic.lean b/src/order/filter/basic.lean index 9b280da249ad0..64c0d6fbd66e5 100644 --- a/src/order/filter/basic.lean +++ b/src/order/filter/basic.lean @@ -2580,8 +2580,9 @@ lemma tendsto.not_tendsto {f : α → β} {a : filter α} {b₁ b₂ : filter β ¬ tendsto f a b₂ := λ hf', (tendsto_inf.2 ⟨hf, hf'⟩).ne_bot.ne hb.eq_bot -lemma tendsto.if {l₁ : filter α} {l₂ : filter β} {f g : α → β} {p : α → Prop} [∀ x, decidable (p x)] - (h₀ : tendsto f (l₁ ⊓ 𝓟 {x | p x}) l₂) (h₁ : tendsto g (l₁ ⊓ 𝓟 { x | ¬ p x }) l₂) : +protected lemma tendsto.if {l₁ : filter α} {l₂ : filter β} {f g : α → β} {p : α → Prop} + [∀ x, decidable (p x)] (h₀ : tendsto f (l₁ ⊓ 𝓟 {x | p x}) l₂) + (h₁ : tendsto g (l₁ ⊓ 𝓟 { x | ¬ p x }) l₂) : tendsto (λ x, if p x then f x else g x) l₁ l₂ := begin simp only [tendsto_def, mem_inf_principal] at *, @@ -2593,8 +2594,8 @@ begin exacts [hp₀ h, hp₁ h], end -lemma tendsto.if' {α β : Type*} {l₁ : filter α} {l₂ : filter β} {f g : α → β} {p : α → Prop} - [decidable_pred p] (hf : tendsto f l₁ l₂) (hg : tendsto g l₁ l₂) : +protected lemma tendsto.if' {α β : Type*} {l₁ : filter α} {l₂ : filter β} {f g : α → β} + {p : α → Prop} [decidable_pred p] (hf : tendsto f l₁ l₂) (hg : tendsto g l₁ l₂) : tendsto (λ a, if p a then f a else g a) l₁ l₂ := begin replace hf : tendsto f (l₁ ⊓ 𝓟 {x | p x}) l₂ := tendsto_inf_left hf, @@ -2602,7 +2603,7 @@ begin exact hf.if hg, end -lemma tendsto.piecewise {l₁ : filter α} {l₂ : filter β} {f g : α → β} +protected lemma tendsto.piecewise {l₁ : filter α} {l₂ : filter β} {f g : α → β} {s : set α} [∀ x, decidable (x ∈ s)] (h₀ : tendsto f (l₁ ⊓ 𝓟 s) l₂) (h₁ : tendsto g (l₁ ⊓ 𝓟 sᶜ) l₂) : tendsto (piecewise s f g) l₁ l₂ := From aa1d645aeadb20d95fc8108d41b89d343f3f503f Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Tue, 31 Jan 2023 16:05:45 +0000 Subject: [PATCH 0394/1291] chore(measure_theory/integrals): API improvements for interval_integrable (#18335) This moves some lemmas which had ended up in a duplicated namespace `interval_integral.interval_integrable`, and fills in a couple of obvious lemmas which were missing. --- src/analysis/special_functions/integrals.lean | 17 +------- .../integral/interval_integral.lean | 39 +++++++++++++++++++ 2 files changed, 40 insertions(+), 16 deletions(-) diff --git a/src/analysis/special_functions/integrals.lean b/src/analysis/special_functions/integrals.lean index f9f502040cfb7..6b6b24c57e75c 100644 --- a/src/analysis/special_functions/integrals.lean +++ b/src/analysis/special_functions/integrals.lean @@ -95,21 +95,6 @@ continuous_id.interval_integrable a b lemma interval_integrable_const : interval_integrable (λ x, c) μ a b := continuous_const.interval_integrable a b -@[simp] -lemma interval_integrable.const_mul (h : interval_integrable f ν a b) : - interval_integrable (λ x, c * f x) ν a b := -by convert h.smul c - -@[simp] -lemma interval_integrable.mul_const (h : interval_integrable f ν a b) : - interval_integrable (λ x, f x * c) ν a b := -by simp only [mul_comm, interval_integrable.const_mul c h] - -@[simp] -lemma interval_integrable.div (h : interval_integrable f ν a b) : - interval_integrable (λ x, f x / c) ν a b := -interval_integrable.mul_const c⁻¹ h - lemma interval_integrable_one_div (h : ∀ x : ℝ, x ∈ [a, b] → f x ≠ 0) (hf : continuous_on f [a, b]) : interval_integrable (λ x, 1 / f x) μ a b := @@ -126,7 +111,7 @@ lemma interval_integrable_exp : interval_integrable exp μ a b := continuous_exp.interval_integrable a b @[simp] -lemma interval_integrable.log +lemma _root_.interval_integrable.log (hf : continuous_on f [a, b]) (h : ∀ x : ℝ, x ∈ [a, b] → f x ≠ 0) : interval_integrable (λ x, log (f x)) μ a b := (continuous_on.log hf h).interval_integrable diff --git a/src/measure_theory/integral/interval_integral.lean b/src/measure_theory/integral/interval_integral.lean index a2f650caf7302..275b56a23c36d 100644 --- a/src/measure_theory/integral/interval_integral.lean +++ b/src/measure_theory/integral/interval_integral.lean @@ -399,14 +399,22 @@ begin exact hf.continuous_on_mul_of_subset hg is_compact_uIcc measurable_set_Ioc Ioc_subset_Icc_self end +@[simp] lemma const_mul {f : ℝ → A} (hf : interval_integrable f μ a b) (c : A) : interval_integrable (λ x, c * f x) μ a b := hf.continuous_on_mul continuous_on_const +@[simp] lemma mul_const {f : ℝ → A} (hf : interval_integrable f μ a b) (c : A) : interval_integrable (λ x, f x * c) μ a b := hf.mul_continuous_on continuous_on_const +@[simp] +lemma div_const {𝕜 : Type*} {f : ℝ → 𝕜} [normed_field 𝕜] + (h : interval_integrable f μ a b) (c : 𝕜) : + interval_integrable (λ x, f x / c) μ a b := +by simpa only [div_eq_mul_inv] using mul_const h c⁻¹ + lemma comp_mul_left (hf : interval_integrable f volume a b) (c : ℝ) : interval_integrable (λ x, f (c * x)) volume (a / c) (b / c) := begin @@ -422,12 +430,43 @@ begin { rw preimage_mul_const_uIcc (inv_ne_zero hc), field_simp [hc] }, end +lemma comp_mul_right (hf : interval_integrable f volume a b) (c : ℝ) : + interval_integrable (λ x, f (x * c)) volume (a / c) (b / c) := +by simpa only [mul_comm] using comp_mul_left hf c + +lemma comp_add_right (hf : interval_integrable f volume a b) (c : ℝ) : + interval_integrable (λ x, f (x + c)) volume (a - c) (b - c) := +begin + wlog h := le_total a b using [a b, b a] tactic.skip, + swap, { exact λ h, interval_integrable.symm (this h.symm) }, + rw interval_integrable_iff' at hf ⊢, + have A : measurable_embedding (λ x, x + c) := + (homeomorph.add_right c).closed_embedding.measurable_embedding, + have Am : measure.map (λ x, x + c) volume = volume, + { exact is_add_left_invariant.is_add_right_invariant.map_add_right_eq_self _ }, + rw ←Am at hf, + convert (measurable_embedding.integrable_on_map_iff A).mp hf, + rw preimage_add_const_uIcc, +end + +lemma comp_add_left (hf : interval_integrable f volume a b) (c : ℝ) : + interval_integrable (λ x, f (c + x)) volume (a - c) (b - c) := +by simpa only [add_comm] using interval_integrable.comp_add_right hf c + +lemma comp_sub_right (hf : interval_integrable f volume a b) (c : ℝ) : + interval_integrable (λ x, f (x - c)) volume (a + c) (b + c) := +by simpa only [sub_neg_eq_add] using interval_integrable.comp_add_right hf (-c) + lemma iff_comp_neg : interval_integrable f volume a b ↔ interval_integrable (λ x, f (-x)) volume (-a) (-b) := begin split, all_goals { intro hf, convert comp_mul_left hf (-1), simp, field_simp, field_simp }, end +lemma comp_sub_left (hf : interval_integrable f volume a b) (c : ℝ) : + interval_integrable (λ x, f (c - x)) volume (c - a) (c - b) := +by simpa only [neg_sub, ←sub_eq_add_neg] using iff_comp_neg.mp (hf.comp_add_left c) + end interval_integrable section From 05a78c9451101108e638a0f213fb1bed82483545 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 31 Jan 2023 16:05:46 +0000 Subject: [PATCH 0395/1291] chore(algebra/char_p/basic): add char_p instances for `ulift` and `mul_opposite` (#18336) --- src/algebra/char_p/basic.lean | 8 +++++++- 1 file changed, 7 insertions(+), 1 deletion(-) diff --git a/src/algebra/char_p/basic.lean b/src/algebra/char_p/basic.lean index 600f16288cfe3..c93798109be1c 100644 --- a/src/algebra/char_p/basic.lean +++ b/src/algebra/char_p/basic.lean @@ -543,7 +543,7 @@ end section prod -variables (S : Type v) [semiring R] [semiring S] (p q : ℕ) [char_p R p] +variables (S : Type v) [add_monoid_with_one R] [add_monoid_with_one S] (p q : ℕ) [char_p R p] /-- The characteristic of the product of rings is the least common multiple of the characteristics of the two rings. -/ @@ -559,6 +559,12 @@ by convert nat.lcm.char_p R S p p; simp end prod +instance ulift.char_p [add_monoid_with_one R] (p : ℕ) [char_p R p] : char_p (ulift.{v} R) p := +{ cast_eq_zero_iff := λ n, iff.trans (ulift.ext_iff _ _) $ char_p.cast_eq_zero_iff R p n } + +instance mul_opposite.char_p [add_monoid_with_one R] (p : ℕ) [char_p R p] : char_p (Rᵐᵒᵖ) p := +{ cast_eq_zero_iff := λ n, mul_opposite.unop_inj.symm.trans $ char_p.cast_eq_zero_iff R p n } + section /-- If two integers from `{0, 1, -1}` result in equal elements in a ring `R` From 59694bd07f0a39c5beccba34bd9f413a160782bf Mon Sep 17 00:00:00 2001 From: sgouezel Date: Tue, 31 Jan 2023 19:36:26 +0000 Subject: [PATCH 0396/1291] feat(measure_theory/integral/peak_function): convergence of integral against a sequence of peak functions (#18327) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit If a sequence of peak functions `φᵢ` converges uniformly to zero away from a point `x₀`, and `g` is integrable and continuous at `x₀`, then `∫ φᵢ • g` converges to `g x₀`. We prove this statement and some consequences of it (notably to sequences which are the successive powers of a given function). Co-authored-by: David Loeffler Co-authored-by: loefflerd --- src/measure_theory/function/l1_space.lean | 15 + src/measure_theory/function/lp_space.lean | 10 + src/measure_theory/integral/bochner.lean | 4 + .../integral/peak_function.lean | 343 ++++++++++++++++++ 4 files changed, 372 insertions(+) create mode 100644 src/measure_theory/integral/peak_function.lean diff --git a/src/measure_theory/function/l1_space.lean b/src/measure_theory/function/l1_space.lean index 7178483822149..062832af9c13a 100644 --- a/src/measure_theory/function/l1_space.lean +++ b/src/measure_theory/function/l1_space.lean @@ -882,6 +882,20 @@ lemma integrable_smul_iff {c : 𝕜} (hc : c ≠ 0) (f : α → β) : integrable (c • f) μ ↔ integrable f μ := and_congr (ae_strongly_measurable_const_smul_iff₀ hc) (has_finite_integral_smul_iff hc f) +lemma integrable.smul_of_top_right {f : α → β} {φ : α → 𝕜} + (hf : integrable f μ) (hφ : mem_ℒp φ ∞ μ) : + integrable (φ • f) μ := +by { rw ← mem_ℒp_one_iff_integrable at hf ⊢, exact mem_ℒp.smul_of_top_right hf hφ } + +lemma integrable.smul_of_top_left {f : α → β} {φ : α → 𝕜} + (hφ : integrable φ μ) (hf : mem_ℒp f ∞ μ) : + integrable (φ • f) μ := +by { rw ← mem_ℒp_one_iff_integrable at hφ ⊢, exact mem_ℒp.smul_of_top_left hf hφ } + +lemma integrable.smul_const {f : α → 𝕜} (hf : integrable f μ) (c : β) : + integrable (λ x, f x • c) μ := +hf.smul_of_top_left (mem_ℒp_top_const c) + end normed_space section normed_space_over_complete_field @@ -897,6 +911,7 @@ begin have : ∀ x : ℝ≥0∞, x = 0 → x < ∞ := by simp, simp [hc, or_iff_left_of_imp (this _)] end + end normed_space_over_complete_field section is_R_or_C diff --git a/src/measure_theory/function/lp_space.lean b/src/measure_theory/function/lp_space.lean index d9985fc0b00c7..f803c1a4891fd 100644 --- a/src/measure_theory/function/lp_space.lean +++ b/src/measure_theory/function/lp_space.lean @@ -1325,6 +1325,16 @@ lemma mem_ℒp.smul {p q r : ℝ≥0∞} {f : α → E} {φ : α → 𝕜} ⟨hφ.1.smul hf.1, (snorm_smul_le_mul_snorm hf.1 hφ.1 hpqr).trans_lt (ennreal.mul_lt_top hφ.snorm_ne_top hf.snorm_ne_top)⟩ +lemma mem_ℒp.smul_of_top_right {p : ℝ≥0∞} {f : α → E} {φ : α → 𝕜} + (hf : mem_ℒp f p μ) (hφ : mem_ℒp φ ∞ μ) : + mem_ℒp (φ • f) p μ := +by { apply hf.smul hφ, simp only [ennreal.div_top, zero_add] } + +lemma mem_ℒp.smul_of_top_left {p : ℝ≥0∞} {f : α → E} {φ : α → 𝕜} + (hf : mem_ℒp f ∞ μ) (hφ : mem_ℒp φ p μ) : + mem_ℒp (φ • f) p μ := +by { apply hf.smul hφ, simp only [ennreal.div_top, add_zero] } + end normed_space section monotonicity diff --git a/src/measure_theory/integral/bochner.lean b/src/measure_theory/integral/bochner.lean index 5d898d490e1bd..e8e570a5d1e8c 100644 --- a/src/measure_theory/integral/bochner.lean +++ b/src/measure_theory/integral/bochner.lean @@ -770,6 +770,10 @@ integral_zero α E variables {α E} +lemma integrable_of_integral_eq_one {f : α → ℝ} (h : ∫ x, f x ∂μ = 1) : + integrable f μ := +by { contrapose h, rw integral_undef h, exact zero_ne_one } + lemma integral_add (hf : integrable f μ) (hg : integrable g μ) : ∫ a, f a + g a ∂μ = ∫ a, f a ∂μ + ∫ a, g a ∂μ := begin diff --git a/src/measure_theory/integral/peak_function.lean b/src/measure_theory/integral/peak_function.lean new file mode 100644 index 0000000000000..ac8e5df9b3b59 --- /dev/null +++ b/src/measure_theory/integral/peak_function.lean @@ -0,0 +1,343 @@ +/- +Copyright (c) 2023 Sébastien Gouëzel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sébastien Gouëzel +-/ +import measure_theory.integral.set_integral +import measure_theory.function.locally_integrable + +/-! +# Integrals against peak functions + +A sequence of peak functions is a sequence of functions with average one concentrating around +a point `x₀`. Given such a sequence `φₙ`, then `∫ φₙ g` tends to `g x₀` in many situations, with +a whole zoo of possible assumptions on `φₙ` and `g`. This file is devoted to such results. + +## Main results + +* `tendsto_set_integral_peak_smul_of_integrable_on_of_continuous_within_at`: If a sequence of peak + functions `φᵢ` converges uniformly to zero away from a point `x₀`, and + `g` is integrable and continuous at `x₀`, then `∫ φᵢ • g` converges to `g x₀`. +* `tendsto_set_integral_pow_smul_of_unique_maximum_of_is_compact_of_continuous_on`: + If a continuous function `c` realizes its maximum at a unique point `x₀` in a compact set `s`, + then the sequence of functions `(c x) ^ n / ∫ (c x) ^ n` is a sequence of peak functions + concentrating around `x₀`. Therefore, `∫ (c x) ^ n * g / ∫ (c x) ^ n` converges to `g x₀` + if `g` is continuous on `s`. + +Note that there are related results about convolution with respect to peak functions in the file +`analysis.convolution`, such as `convolution_tendsto_right` there. +-/ + +open set filter measure_theory measure_theory.measure topological_space metric +open_locale topology ennreal + +/-- This lemma exists for finsets, but not for sets currently. porting note: move to +data.set.basic after the port. -/ +lemma set.disjoint_sdiff_inter {α : Type*} (s t : set α) : disjoint (s \ t) (s ∩ t) := +disjoint_of_subset_right (inter_subset_right _ _) disjoint_sdiff_left + +open set + +variables {α E ι : Type*} {hm : measurable_space α} {μ : measure α} + [topological_space α] [borel_space α] + [normed_add_comm_group E] [normed_space ℝ E] + {g : α → E} {l : filter ι} {x₀ : α} {s : set α} {φ : ι → α → ℝ} + +/-- If a sequence of peak functions `φᵢ` converges uniformly to zero away from a point `x₀`, and +`g` is integrable and continuous at `x₀`, then `φᵢ • g` is eventually integrable. -/ +lemma integrable_on_peak_smul_of_integrable_on_of_continuous_within_at + (hs : measurable_set s) + (hlφ : ∀ (u : set α), is_open u → x₀ ∈ u → tendsto_uniformly_on φ 0 l (s \ u)) + (hiφ : ∀ᶠ i in l, ∫ x in s, φ i x ∂μ = 1) + (hmg : integrable_on g s μ) + (hcg : continuous_within_at g s x₀) : + ∀ᶠ i in l, integrable_on (λ x, φ i x • g x) s μ := +begin + obtain ⟨u, u_open, x₀u, hu⟩ : ∃ u, is_open u ∧ x₀ ∈ u ∧ ∀ x ∈ u ∩ s, g x ∈ ball (g x₀) 1, + from mem_nhds_within.1 (hcg (ball_mem_nhds _ zero_lt_one)), + filter_upwards [tendsto_uniformly_on_iff.1 (hlφ u u_open x₀u) 1 zero_lt_one, hiφ] + with i hi h'i, + have A : integrable_on (λ x, φ i x • g x) (s \ u) μ, + { apply integrable.smul_of_top_right (hmg.mono (diff_subset _ _) le_rfl), + apply mem_ℒp_top_of_bound + ((integrable_of_integral_eq_one h'i).ae_strongly_measurable.mono_set ((diff_subset _ _))) 1, + filter_upwards [self_mem_ae_restrict (hs.diff u_open.measurable_set)] with x hx, + simpa only [pi.zero_apply, dist_zero_left] using (hi x hx).le }, + have B : integrable_on (λ x, φ i x • g x) (s ∩ u) μ, + { apply integrable.smul_of_top_left, + { exact integrable_on.mono_set (integrable_of_integral_eq_one h'i) (inter_subset_left _ _) }, + { apply mem_ℒp_top_of_bound (hmg.mono_set (inter_subset_left _ _)).ae_strongly_measurable + (‖g x₀‖ + 1), + filter_upwards [self_mem_ae_restrict (hs.inter u_open.measurable_set)] with x hx, + rw inter_comm at hx, + exact (norm_lt_of_mem_ball (hu x hx)).le } }, + convert A.union B, + simp only [diff_union_inter], +end + +variables [complete_space E] + +/-- If a sequence of peak functions `φᵢ` converges uniformly to zero away from a point `x₀`, and +`g` is integrable and continuous at `x₀`, then `∫ φᵢ • g` converges to `x₀`. Auxiliary lemma +where one assumes additionally `g x₀ = 0`. -/ +lemma tendsto_set_integral_peak_smul_of_integrable_on_of_continuous_within_at_aux + (hs : measurable_set s) + (hnφ : ∀ᶠ i in l, ∀ x ∈ s, 0 ≤ φ i x) + (hlφ : ∀ (u : set α), is_open u → x₀ ∈ u → tendsto_uniformly_on φ 0 l (s \ u)) + (hiφ : ∀ᶠ i in l, ∫ x in s, φ i x ∂μ = 1) + (hmg : integrable_on g s μ) (h'g : g x₀ = 0) + (hcg : continuous_within_at g s x₀) : + tendsto (λ i : ι, ∫ x in s, φ i x • g x ∂μ) l (𝓝 0) := +begin + refine metric.tendsto_nhds.2 (λ ε εpos, _), + obtain ⟨δ, hδ, δpos⟩ : ∃ δ, δ * ∫ x in s, ‖g x‖ ∂μ + δ < ε ∧ 0 < δ, + { have A : tendsto (λ δ, δ * ∫ x in s, ‖g x‖ ∂μ + δ) (𝓝[>] 0) (𝓝 (0 * ∫ x in s, ‖g x‖ ∂μ + 0)), + { apply tendsto.mono_left _ nhds_within_le_nhds, + exact (tendsto_id.mul tendsto_const_nhds).add tendsto_id }, + rw [zero_mul, zero_add] at A, + exact (((tendsto_order.1 A).2 ε εpos).and self_mem_nhds_within).exists }, + suffices : ∀ᶠ i in l, ‖∫ x in s, φ i x • g x ∂μ‖ ≤ δ * ∫ x in s, ‖g x‖ ∂μ + δ, + { filter_upwards [this] with i hi, + simp only [dist_zero_right], + exact hi.trans_lt hδ }, + obtain ⟨u, u_open, x₀u, hu⟩ : ∃ u, is_open u ∧ x₀ ∈ u ∧ ∀ x ∈ u ∩ s, g x ∈ ball (g x₀) δ, + from mem_nhds_within.1 (hcg (ball_mem_nhds _ δpos)), + filter_upwards [tendsto_uniformly_on_iff.1 (hlφ u u_open x₀u) δ δpos, hiφ, hnφ, + integrable_on_peak_smul_of_integrable_on_of_continuous_within_at hs hlφ hiφ hmg hcg] + with i hi h'i hφpos h''i, + have B : ‖∫ x in s ∩ u, φ i x • g x ∂μ‖ ≤ δ, from calc + ‖∫ x in s ∩ u, φ i x • g x ∂μ‖ ≤ ∫ x in s ∩ u, ‖φ i x • g x‖ ∂μ : + norm_integral_le_integral_norm _ + ... ≤ ∫ x in s ∩ u, ‖φ i x‖ * δ ∂μ : + begin + refine set_integral_mono_on _ _ (hs.inter u_open.measurable_set) (λ x hx, _), + { exact integrable_on.mono_set h''i.norm (inter_subset_left _ _) }, + { exact integrable_on.mono_set ((integrable_of_integral_eq_one h'i).norm.mul_const _) + (inter_subset_left _ _) }, + rw norm_smul, + apply mul_le_mul_of_nonneg_left _ (norm_nonneg _), + rw [inter_comm, h'g] at hu, + exact (mem_ball_zero_iff.1 (hu x hx)).le, + end + ... ≤ ∫ x in s, ‖φ i x‖ * δ ∂μ : + begin + apply set_integral_mono_set, + { exact ((integrable_of_integral_eq_one h'i).norm.mul_const _) }, + { exact eventually_of_forall (λ x, mul_nonneg (norm_nonneg _) δpos.le) }, + { apply eventually_of_forall, exact inter_subset_left s u } + end + ... = ∫ x in s, φ i x * δ ∂μ : + begin + apply set_integral_congr hs (λ x hx, _), + rw real.norm_of_nonneg (hφpos _ hx), + end + ... = δ : by rw [integral_mul_right, h'i, one_mul], + have C : ‖∫ x in s \ u, φ i x • g x ∂μ‖ ≤ δ * ∫ x in s, ‖g x‖ ∂μ, from calc + ‖∫ x in s \ u, φ i x • g x ∂μ‖ ≤ ∫ x in s \ u, ‖φ i x • g x‖ ∂μ : + norm_integral_le_integral_norm _ + ... ≤ ∫ x in s \ u, δ * ‖g x‖ ∂μ : + begin + refine set_integral_mono_on _ _ (hs.diff u_open.measurable_set) (λ x hx, _), + { exact integrable_on.mono_set h''i.norm (diff_subset _ _) }, + { exact integrable_on.mono_set (hmg.norm.const_mul _) (diff_subset _ _) }, + rw norm_smul, + apply mul_le_mul_of_nonneg_right _ (norm_nonneg _), + simpa only [pi.zero_apply, dist_zero_left] using (hi x hx).le, + end + ... ≤ δ * ∫ x in s, ‖g x‖ ∂μ : + begin + rw integral_mul_left, + apply mul_le_mul_of_nonneg_left (set_integral_mono_set hmg.norm _ _) δpos.le, + { exact eventually_of_forall (λ x, norm_nonneg _) }, + { apply eventually_of_forall, exact diff_subset s u } + end, + calc + ‖∫ x in s, φ i x • g x ∂μ‖ = ‖∫ x in s \ u, φ i x • g x ∂μ + ∫ x in s ∩ u, φ i x • g x ∂μ‖ : + begin + conv_lhs { rw ← diff_union_inter s u }, + rw integral_union (disjoint_sdiff_inter _ _) (hs.inter u_open.measurable_set) + (h''i.mono_set (diff_subset _ _)) (h''i.mono_set (inter_subset_left _ _)) + end + ... ≤ ‖∫ x in s \ u, φ i x • g x ∂μ‖ + ‖∫ x in s ∩ u, φ i x • g x ∂μ‖ : norm_add_le _ _ + ... ≤ δ * ∫ x in s, ‖g x‖ ∂μ + δ : add_le_add C B +end + +/- If a sequence of peak functions `φᵢ` converges uniformly to zero away from a point `x₀`, and +`g` is integrable and continuous at `x₀`, then `∫ φᵢ • g` converges to `x₀`. -/ +lemma tendsto_set_integral_peak_smul_of_integrable_on_of_continuous_within_at + (hs : measurable_set s) (h's : μ s ≠ ∞) + (hnφ : ∀ᶠ i in l, ∀ x ∈ s, 0 ≤ φ i x) + (hlφ : ∀ (u : set α), is_open u → x₀ ∈ u → tendsto_uniformly_on φ 0 l (s \ u)) + (hiφ : (λ i, ∫ x in s, φ i x ∂μ) =ᶠ[l] 1) + (hmg : integrable_on g s μ) + (hcg : continuous_within_at g s x₀) : + tendsto (λ i : ι, ∫ x in s, φ i x • g x ∂μ) l (𝓝 (g x₀)) := +begin + let h := g - (λ y, g x₀), + have A : tendsto (λ i : ι, ∫ x in s, φ i x • h x ∂μ + (∫ x in s, φ i x ∂μ) • g x₀) l + (𝓝 (0 + (1 : ℝ) • g x₀)), + { refine tendsto.add _ (tendsto.smul (tendsto_const_nhds.congr' hiφ.symm) tendsto_const_nhds), + apply tendsto_set_integral_peak_smul_of_integrable_on_of_continuous_within_at_aux + hs hnφ hlφ hiφ, + { apply integrable.sub hmg, + apply integrable_on_const.2, + simp only [h's.lt_top, or_true] }, + { simp only [h, pi.sub_apply, sub_self] }, + { exact hcg.sub continuous_within_at_const } }, + simp only [one_smul, zero_add] at A, + refine tendsto.congr' _ A, + filter_upwards [integrable_on_peak_smul_of_integrable_on_of_continuous_within_at + hs hlφ hiφ hmg hcg, hiφ] with i hi h'i, + simp only [h, pi.sub_apply, smul_sub], + rw [integral_sub hi, integral_smul_const, sub_add_cancel], + exact integrable.smul_const (integrable_of_integral_eq_one h'i) _, +end + +/-- If a continuous function `c` realizes its maximum at a unique point `x₀` in a compact set `s`, +then the sequence of functions `(c x) ^ n / ∫ (c x) ^ n` is a sequence of peak functions +concentrating around `x₀`. Therefore, `∫ (c x) ^ n * g / ∫ (c x) ^ n` converges to `g x₀` if `g` is +integrable on `s` and continuous at `x₀`. + +Version assuming that `μ` gives positive mass to all neighborhoods of `x₀` within `s`. +For a less precise but more usable version, see +`tendsto_set_integral_pow_smul_of_unique_maximum_of_is_compact_of_continuous_on`. + -/ +lemma tendsto_set_integral_pow_smul_of_unique_maximum_of_is_compact_of_measure_nhds_within_pos + [metrizable_space α] [is_locally_finite_measure μ] (hs : is_compact s) + (hμ : ∀ u, is_open u → x₀ ∈ u → 0 < μ (u ∩ s)) + {c : α → ℝ} (hc : continuous_on c s) (h'c : ∀ y ∈ s, y ≠ x₀ → c y < c x₀) + (hnc : ∀ x ∈ s, 0 ≤ c x) (hnc₀ : 0 < c x₀) (h₀ : x₀ ∈ s) + (hmg : integrable_on g s μ) + (hcg : continuous_within_at g s x₀) : + tendsto (λ (n : ℕ), (∫ x in s, (c x) ^ n ∂μ)⁻¹ • (∫ x in s, (c x) ^ n • g x ∂μ)) at_top + (𝓝 (g x₀)) := +begin + /- We apply the general result + `tendsto_set_integral_peak_smul_of_integrable_on_of_continuous_within_at` to the sequence of + peak functions `φₙ = (c x) ^ n / ∫ (c x) ^ n`. The only nontrivial bit is to check that this + sequence converges uniformly to zero on any set `s \ u` away from `x₀`. By compactness, the + function `c` is bounded by `t < c x₀` there. Consider `t' ∈ (t, c x₀)`, and a neighborhood `v` + of `x₀` where `c x ≥ t'`, by continuity. Then `∫ (c x) ^ n` is bounded below by `t' ^ n μ v`. + It follows that, on `s \ u`, then `φₙ x ≤ t ^ n / (t' ^ n μ v)`, which tends (exponentially fast) + to zero with `n`. -/ + let φ : ℕ → α → ℝ := λ n x, (∫ x in s, (c x) ^ n ∂μ)⁻¹ * (c x) ^ n, + have hnφ : ∀ n, ∀ x ∈ s, 0 ≤ φ n x, + { assume n x hx, + apply mul_nonneg (inv_nonneg.2 _) (pow_nonneg (hnc x hx) _), + exact set_integral_nonneg hs.measurable_set (λ x hx, pow_nonneg (hnc x hx) _) }, + have I : ∀ n, integrable_on (λ x, (c x)^n) s μ := + λ n, continuous_on.integrable_on_compact hs (hc.pow n), + have J : ∀ n, 0 ≤ᵐ[μ.restrict s] λ (x : α), c x ^ n, + { assume n, + filter_upwards [ae_restrict_mem hs.measurable_set] with x hx, + exact pow_nonneg (hnc x hx) n }, + have P : ∀ n, 0 < ∫ x in s, (c x) ^ n ∂μ, + { assume n, + refine (set_integral_pos_iff_support_of_nonneg_ae (J n) (I n)).2 _, + obtain ⟨u, u_open, x₀_u, hu⟩ : ∃ (u : set α), is_open u ∧ x₀ ∈ u ∧ u ∩ s ⊆ c ⁻¹' Ioi 0 := + _root_.continuous_on_iff.1 hc x₀ h₀ (Ioi (0 : ℝ)) is_open_Ioi hnc₀, + apply (hμ u u_open x₀_u).trans_le, + exact measure_mono (λ x hx, ⟨ne_of_gt (pow_pos (hu hx) _), hx.2⟩) }, + have hiφ : ∀ n, ∫ x in s, φ n x ∂μ = 1 := + λ n, by rw [integral_mul_left, inv_mul_cancel (P n).ne'], + have A : ∀ (u : set α), is_open u → x₀ ∈ u → tendsto_uniformly_on φ 0 at_top (s \ u), + { assume u u_open x₀u, + obtain ⟨t, t_pos, tx₀, ht⟩ : ∃ t, 0 ≤ t ∧ t < c x₀ ∧ (∀ x ∈ s \ u, c x ≤ t), + { rcases eq_empty_or_nonempty (s \ u) with h|h, + { exact ⟨0, le_rfl, hnc₀, + by simp only [h, mem_empty_iff_false, is_empty.forall_iff, implies_true_iff]⟩ }, + obtain ⟨x, hx, h'x⟩ : ∃ x ∈ s \ u, ∀ y ∈ s \ u, c y ≤ c x := + is_compact.exists_forall_ge (hs.diff u_open) h (hc.mono (diff_subset _ _)), + refine ⟨c x, hnc x hx.1, h'c x hx.1 _, h'x⟩, + rintros rfl, + exact hx.2 x₀u }, + obtain ⟨t', tt', t'x₀⟩ : ∃ t', t < t' ∧ t' < c x₀ := exists_between tx₀, + have t'_pos : 0 < t' := t_pos.trans_lt tt', + obtain ⟨v, v_open, x₀_v, hv⟩ : ∃ (v : set α), is_open v ∧ x₀ ∈ v ∧ v ∩ s ⊆ c ⁻¹' Ioi t' := + _root_.continuous_on_iff.1 hc x₀ h₀ (Ioi t') is_open_Ioi t'x₀, + have M : ∀ n, ∀ x ∈ s \ u, φ n x ≤ (μ (v ∩ s)).to_real ⁻¹ * (t / t') ^ n, + { assume n x hx, + have B : t' ^ n * (μ (v ∩ s)).to_real ≤ ∫ y in s, (c y) ^ n ∂μ, from calc + t' ^ n * (μ (v ∩ s)).to_real = ∫ y in v ∩ s, t' ^ n ∂μ : + by simp only [integral_const, measure.restrict_apply, measurable_set.univ, univ_inter, + algebra.id.smul_eq_mul, mul_comm] + ... ≤ ∫ y in v ∩ s, (c y) ^ n ∂μ : + begin + apply set_integral_mono_on _ _ (v_open.measurable_set.inter hs.measurable_set) _, + { apply integrable_on_const.2 (or.inr _), + exact lt_of_le_of_lt (measure_mono (inter_subset_right _ _)) hs.measure_lt_top }, + { exact (I n).mono (inter_subset_right _ _) le_rfl }, + { assume x hx, + exact pow_le_pow_of_le_left t'_pos.le (le_of_lt (hv hx)) _ } + end + ... ≤ ∫ y in s, (c y) ^ n ∂μ : + set_integral_mono_set (I n) (J n) (eventually_of_forall (inter_subset_right _ _)), + simp_rw [φ, ← div_eq_inv_mul, div_pow, div_div], + apply div_le_div (pow_nonneg t_pos n) _ _ B, + { exact pow_le_pow_of_le_left (hnc _ hx.1) (ht x hx) _ }, + { apply mul_pos (pow_pos (t_pos.trans_lt tt') _) + (ennreal.to_real_pos (hμ v v_open x₀_v).ne' _), + have : μ (v ∩ s) ≤ μ s := measure_mono (inter_subset_right _ _), + exact ne_of_lt (lt_of_le_of_lt this hs.measure_lt_top) } }, + have N : tendsto (λ n, (μ (v ∩ s)).to_real ⁻¹ * (t / t') ^ n) at_top + (𝓝 ((μ (v ∩ s)).to_real ⁻¹ * 0)), + { apply tendsto.mul tendsto_const_nhds _, { apply_instance }, + apply tendsto_pow_at_top_nhds_0_of_lt_1 (div_nonneg t_pos t'_pos.le), + exact (div_lt_one t'_pos).2 tt' }, + rw mul_zero at N, + refine tendsto_uniformly_on_iff.2 (λ ε εpos, _), + filter_upwards [(tendsto_order.1 N).2 ε εpos] with n hn x hx, + simp only [pi.zero_apply, dist_zero_left, real.norm_of_nonneg (hnφ n x hx.1)], + exact (M n x hx).trans_lt hn }, + have : tendsto (λ (i : ℕ), ∫ (x : α) in s, φ i x • g x ∂μ) at_top (𝓝 (g x₀)) := + tendsto_set_integral_peak_smul_of_integrable_on_of_continuous_within_at hs.measurable_set + hs.measure_lt_top.ne (eventually_of_forall hnφ) A (eventually_of_forall hiφ) hmg hcg, + convert this, + simp_rw [← smul_smul, integral_smul], +end + +/-- If a continuous function `c` realizes its maximum at a unique point `x₀` in a compact set `s`, +then the sequence of functions `(c x) ^ n / ∫ (c x) ^ n` is a sequence of peak functions +concentrating around `x₀`. Therefore, `∫ (c x) ^ n * g / ∫ (c x) ^ n` converges to `g x₀` if `g` is +integrable on `s` and continuous at `x₀`. + +Version assuming that `μ` gives positive mass to all open sets. +For a less precise but more usable version, see +`tendsto_set_integral_pow_smul_of_unique_maximum_of_is_compact_of_continuous_on`. +-/ +lemma tendsto_set_integral_pow_smul_of_unique_maximum_of_is_compact_of_integrable_on + [metrizable_space α] [is_locally_finite_measure μ] [is_open_pos_measure μ] (hs : is_compact s) + {c : α → ℝ} (hc : continuous_on c s) (h'c : ∀ y ∈ s, y ≠ x₀ → c y < c x₀) + (hnc : ∀ x ∈ s, 0 ≤ c x) (hnc₀ : 0 < c x₀) (h₀ : x₀ ∈ closure (interior s)) + (hmg : integrable_on g s μ) + (hcg : continuous_within_at g s x₀) : + tendsto (λ (n : ℕ), (∫ x in s, (c x) ^ n ∂μ)⁻¹ • (∫ x in s, (c x) ^ n • g x ∂μ)) at_top + (𝓝 (g x₀)) := +begin + have : x₀ ∈ s, + { rw ← hs.is_closed.closure_eq, exact closure_mono interior_subset h₀ }, + apply tendsto_set_integral_pow_smul_of_unique_maximum_of_is_compact_of_measure_nhds_within_pos + hs _ hc h'c hnc hnc₀ this hmg hcg, + assume u u_open x₀_u, + calc 0 < μ (u ∩ interior s) : + (u_open.inter is_open_interior).measure_pos μ (_root_.mem_closure_iff.1 h₀ u u_open x₀_u) + ... ≤ μ (u ∩ s) : measure_mono (inter_subset_inter_right _ interior_subset) +end + +/-- If a continuous function `c` realizes its maximum at a unique point `x₀` in a compact set `s`, +then the sequence of functions `(c x) ^ n / ∫ (c x) ^ n` is a sequence of peak functions +concentrating around `x₀`. Therefore, `∫ (c x) ^ n * g / ∫ (c x) ^ n` converges to `g x₀` if `g` is +continuous on `s`. -/ +lemma tendsto_set_integral_pow_smul_of_unique_maximum_of_is_compact_of_continuous_on + [metrizable_space α] [is_locally_finite_measure μ] [is_open_pos_measure μ] (hs : is_compact s) + {c : α → ℝ} (hc : continuous_on c s) (h'c : ∀ y ∈ s, y ≠ x₀ → c y < c x₀) + (hnc : ∀ x ∈ s, 0 ≤ c x) (hnc₀ : 0 < c x₀) (h₀ : x₀ ∈ closure (interior s)) + (hmg : continuous_on g s) : + tendsto (λ (n : ℕ), (∫ x in s, (c x) ^ n ∂μ)⁻¹ • (∫ x in s, (c x) ^ n • g x ∂μ)) at_top + (𝓝 (g x₀)) := +begin + have : x₀ ∈ s, + { rw ← hs.is_closed.closure_eq, exact closure_mono interior_subset h₀ }, + exact tendsto_set_integral_pow_smul_of_unique_maximum_of_is_compact_of_integrable_on + hs hc h'c hnc hnc₀ h₀ (hmg.integrable_on_compact hs) (hmg x₀ this) +end From 832a8ba8f10f11fea99367c469ff802e69a5b8ec Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Tue, 31 Jan 2023 21:08:57 +0000 Subject: [PATCH 0397/1291] feat(topology/metric_space/isometric_smul): new file (#18130) Define typeclasses `has_isometric_smul` and `has_isometric_vadd`. To minimize diff size, I did not fully deduplicate lemmas in this PR. --- .../complex/upper_half_plane/metric.lean | 2 +- src/analysis/normed/group/add_torsor.lean | 37 +- src/analysis/normed/group/basic.lean | 96 +---- src/analysis/normed_space/basic.lean | 8 +- src/analysis/normed_space/pointwise.lean | 14 +- src/analysis/seminorm.lean | 4 +- src/measure_theory/measure/haar_lebesgue.lean | 4 +- src/topology/metric_space/isometric_smul.lean | 395 ++++++++++++++++++ src/topology/metric_space/isometry.lean | 11 +- 9 files changed, 433 insertions(+), 138 deletions(-) create mode 100644 src/topology/metric_space/isometric_smul.lean diff --git a/src/analysis/complex/upper_half_plane/metric.lean b/src/analysis/complex/upper_half_plane/metric.lean index 2811716915ee1..8ad5e99cc6fde 100644 --- a/src/analysis/complex/upper_half_plane/metric.lean +++ b/src/analysis/complex/upper_half_plane/metric.lean @@ -344,7 +344,7 @@ lemma isometry_pos_mul (a : {x : ℝ // 0 < x}) : isometry ((•) a : ℍ → begin refine isometry.of_dist_eq (λ y₁ y₂, _), simp only [dist_eq, coe_pos_real_smul, pos_real_im], congr' 2, - rw [dist_smul, mul_mul_mul_comm, real.sqrt_mul (mul_self_nonneg _), real.sqrt_mul_self_eq_abs, + rw [dist_smul₀, mul_mul_mul_comm, real.sqrt_mul (mul_self_nonneg _), real.sqrt_mul_self_eq_abs, real.norm_eq_abs, mul_left_comm], exact mul_div_mul_left _ _ (mt _root_.abs_eq_zero.1 a.2.ne') end diff --git a/src/analysis/normed/group/add_torsor.lean b/src/analysis/normed/group/add_torsor.lean index e72f7cbeeb169..4a526f6fa62a1 100644 --- a/src/analysis/normed/group/add_torsor.lean +++ b/src/analysis/normed/group/add_torsor.lean @@ -38,6 +38,10 @@ instance normed_add_torsor.to_add_torsor' {V P : Type*} [normed_add_comm_group V variables {α V P W Q : Type*} [seminormed_add_comm_group V] [pseudo_metric_space P] [normed_add_torsor V P] [normed_add_comm_group W] [metric_space Q] [normed_add_torsor W Q] +@[priority 100] +instance normed_add_torsor.to_has_isometric_vadd : has_isometric_vadd V P := +⟨λ c, isometry.of_dist_eq $ λ x y, by simp [normed_add_torsor.dist_eq_norm']⟩ + /-- A `seminormed_add_comm_group` is a `normed_add_torsor` over itself. -/ @[priority 100] instance seminormed_add_comm_group.to_normed_add_torsor : normed_add_torsor V V := @@ -69,9 +73,9 @@ lemma dist_eq_norm_vsub' (x y : P) : dist x y = ‖y -ᵥ x‖ := end -@[simp] lemma dist_vadd_cancel_left (v : V) (x y : P) : +lemma dist_vadd_cancel_left (v : V) (x y : P) : dist (v +ᵥ x) (v +ᵥ y) = dist x y := -by rw [dist_eq_norm_vsub V, dist_eq_norm_vsub V, vadd_vsub_vadd_cancel_left] +dist_vadd _ _ _ @[simp] lemma dist_vadd_cancel_right (v₁ v₂ : V) (x : P) : dist (v₁ +ᵥ x) (v₂ +ᵥ x) = dist v₁ v₂ := @@ -89,17 +93,6 @@ addition/subtraction of `x : P`. -/ { to_equiv := equiv.vadd_const x, isometry_to_fun := isometry.of_dist_eq $ λ _ _, dist_vadd_cancel_right _ _ _ } -section - -variable (P) - -/-- Self-isometry of a (semi)normed add torsor given by addition of a constant vector `x`. -/ -@[simps] def isometry_equiv.const_vadd (x : V) : P ≃ᵢ P := -{ to_equiv := equiv.const_vadd P x, - isometry_to_fun := isometry.of_dist_eq $ λ _ _, dist_vadd_cancel_left _ _ _ } - -end - @[simp] lemma dist_vsub_cancel_left (x y z : P) : dist (x -ᵥ y) (x -ᵥ z) = dist y z := by rw [dist_eq_norm, vsub_sub_vsub_cancel_left, dist_comm, dist_eq_norm_vsub V] @@ -112,24 +105,6 @@ subtraction from `x : P`. -/ @[simp] lemma dist_vsub_cancel_right (x y z : P) : dist (x -ᵥ z) (y -ᵥ z) = dist x y := (isometry_equiv.vadd_const z).symm.dist_eq x y -section pointwise - -open_locale pointwise - -@[simp] lemma vadd_ball (x : V) (y : P) (r : ℝ) : - x +ᵥ metric.ball y r = metric.ball (x +ᵥ y) r := -(isometry_equiv.const_vadd P x).image_ball y r - -@[simp] lemma vadd_closed_ball (x : V) (y : P) (r : ℝ) : - x +ᵥ metric.closed_ball y r = metric.closed_ball (x +ᵥ y) r := -(isometry_equiv.const_vadd P x).image_closed_ball y r - -@[simp] lemma vadd_sphere (x : V) (y : P) (r : ℝ) : - x +ᵥ metric.sphere y r = metric.sphere (x +ᵥ y) r := -(isometry_equiv.const_vadd P x).image_sphere y r - -end pointwise - lemma dist_vadd_vadd_le (v v' : V) (p p' : P) : dist (v +ᵥ p) (v' +ᵥ p') ≤ dist v v' + dist p p' := by simpa using dist_triangle (v +ᵥ p) (v' +ᵥ p) (v' +ᵥ p') diff --git a/src/analysis/normed/group/basic.lean b/src/analysis/normed/group/basic.lean index feb287179b714..0579e29219a64 100644 --- a/src/analysis/normed/group/basic.lean +++ b/src/analysis/normed/group/basic.lean @@ -8,7 +8,7 @@ import order.liminf_limsup import topology.algebra.uniform_group import topology.instances.rat import topology.metric_space.algebra -import topology.metric_space.isometry +import topology.metric_space.isometric_smul import topology.sequences /-! @@ -292,6 +292,9 @@ lemma dist_eq_norm_div' (a b : E) : dist a b = ‖b / a‖ := by rw [dist_comm, alias dist_eq_norm_sub ← dist_eq_norm alias dist_eq_norm_sub' ← dist_eq_norm' +@[to_additive] instance normed_group.to_has_isometric_smul_right : has_isometric_smul Eᵐᵒᵖ E := +⟨λ a, isometry.of_dist_eq $ λ b c, by simp [dist_eq_norm_div]⟩ + @[simp, to_additive] lemma dist_one_right (a : E) : dist a 1 = ‖a‖ := by rw [dist_eq_norm_div, div_one] @@ -313,18 +316,12 @@ by simpa only [dist_eq_norm_div] using dist_comm a b @[simp, to_additive norm_neg] lemma norm_inv' (a : E) : ‖a⁻¹‖ = ‖a‖ := by simpa using norm_div_rev 1 a -@[simp, to_additive] lemma dist_mul_right (a₁ a₂ b : E) : dist (a₁ * b) (a₂ * b) = dist a₁ a₂ := -by simp [dist_eq_norm_div] - @[simp, to_additive] lemma dist_mul_self_right (a b : E) : dist b (a * b) = ‖a‖ := by rw [←dist_one_left, ←dist_mul_right 1 a b, one_mul] @[simp, to_additive] lemma dist_mul_self_left (a b : E) : dist (a * b) b = ‖a‖ := by rw [dist_comm, dist_mul_self_right] -@[to_additive] lemma dist_div_right (a₁ a₂ b : E) : dist (a₁ / b) (a₂ / b) = dist a₁ a₂ := -by simpa only [div_eq_mul_inv] using dist_mul_right _ _ _ - @[simp, to_additive] lemma dist_div_eq_dist_mul_left (a b c : E) : dist (a / b) c = dist a (c * b) := by rw [←dist_mul_right _ _ b, div_mul_cancel'] @@ -494,31 +491,6 @@ def norm_group_seminorm : group_seminorm E := ⟨norm, norm_one', norm_mul_le', variables {E} -namespace isometry_equiv --- TODO This material is superseded by similar constructions such as --- `affine_isometry_equiv.const_vadd`; deduplicate - -/-- Multiplication `y ↦ y * x` as an `isometry`. -/ -@[to_additive "Addition `y ↦ y + x` as an `isometry`"] -protected def mul_right (x : E) : E ≃ᵢ E := -{ isometry_to_fun := isometry.of_dist_eq $ λ y z, dist_mul_right _ _ _, - .. equiv.mul_right x } - -@[simp, to_additive] -lemma mul_right_to_equiv (x : E) : (isometry_equiv.mul_right x).to_equiv = equiv.mul_right x := rfl - -@[simp, to_additive] -lemma coe_mul_right (x : E) : (isometry_equiv.mul_right x : E → E) = λ y, y * x := rfl - -@[to_additive] lemma mul_right_apply (x y : E) : (isometry_equiv.mul_right x : E → E) y = y * x := -rfl - -@[simp, to_additive] -lemma mul_right_symm (x : E) : (isometry_equiv.mul_right x).symm = isometry_equiv.mul_right x⁻¹ := -ext $ λ y, rfl - -end isometry_equiv - @[to_additive] lemma normed_comm_group.tendsto_nhds_one {f : α → E} {l : filter α} : tendsto f l (𝓝 1) ↔ ∀ ε > 0, ∀ᶠ x in l, ‖ f x ‖ < ε := metric.tendsto_nhds.trans $ by simp only [dist_one_right] @@ -672,12 +644,6 @@ by rw [edist_eq_coe_nnnorm_div, div_one] lemma mem_emetric_ball_one_iff {r : ℝ≥0∞} : a ∈ emetric.ball (1 : E) r ↔ ↑‖a‖₊ < r := by rw [emetric.mem_ball, edist_eq_coe_nnnorm'] -@[simp, to_additive] lemma edist_mul_right (a₁ a₂ b : E) : edist (a₁ * b) (a₂ * b) = edist a₁ a₂ := -by simp [edist_dist] - -@[simp, to_additive] lemma edist_div_right (a₁ a₂ b : E) : edist (a₁ / b) (a₂ / b) = edist a₁ a₂ := -by simpa only [div_eq_mul_inv] using edist_mul_right _ _ _ - @[to_additive] lemma monoid_hom_class.lipschitz_of_bound_nnnorm [monoid_hom_class 𝓕 E F] (f : 𝓕) (C : ℝ≥0) (h : ∀ x, ‖f x‖₊ ≤ C * ‖x‖₊) : lipschitz_with C f := @real.to_nnreal_coe C ▸ monoid_hom_class.lipschitz_of_bound f C h @@ -960,18 +926,12 @@ end induced section seminormed_comm_group variables [seminormed_comm_group E] [seminormed_comm_group F] {a a₁ a₂ b b₁ b₂ : E} {r r₁ r₂ : ℝ} -@[simp, to_additive] lemma dist_mul_left (a b₁ b₂ : E) : dist (a * b₁) (a * b₂) = dist b₁ b₂ := -by simp [dist_eq_norm_div] +@[to_additive] instance normed_group.to_has_isometric_smul_left : has_isometric_smul E E := +⟨λ a, isometry.of_dist_eq $ λ b c, by simp [dist_eq_norm_div]⟩ @[to_additive] lemma dist_inv (x y : E) : dist x⁻¹ y = dist x y⁻¹ := by simp_rw [dist_eq_norm_div, ←norm_inv' (x⁻¹ / y), inv_div, div_inv_eq_mul, mul_comm] -@[simp, to_additive] lemma dist_inv_inv (a b : E) : dist a⁻¹ b⁻¹ = dist a b := -by rw [dist_inv, inv_inv] - -@[simp, to_additive] lemma dist_div_left (a b₁ b₂ : E) : dist (a / b₁) (a / b₂) = dist b₁ b₂ := -by simp only [div_eq_mul_inv, dist_mul_left, dist_inv_inv] - @[simp, to_additive] lemma dist_self_mul_right (a b : E) : dist a (a * b) = ‖b‖ := by rw [←dist_one_left, ←dist_mul_left a 1 b, mul_one] @@ -1109,38 +1069,6 @@ by { ext, simp [mem_closed_ball, set.mem_smul_set, dist_eq_norm_div, div_eq_inv_ by { ext, simp [mem_ball, set.mem_smul_set, dist_eq_norm_div, div_eq_inv_mul, ← eq_inv_mul_iff_mul_eq, mul_assoc], } -namespace isometry_equiv - -/-- Multiplication `y ↦ x * y` as an `isometry`. -/ -@[to_additive "Addition `y ↦ x + y` as an `isometry`"] -protected def mul_left (x : E) : E ≃ᵢ E := -{ isometry_to_fun := isometry.of_dist_eq $ λ y z, dist_mul_left _ _ _, - to_equiv := equiv.mul_left x } - -@[simp, to_additive] lemma mul_left_to_equiv (x : E) : - (isometry_equiv.mul_left x).to_equiv = equiv.mul_left x := rfl - -@[simp, to_additive] lemma coe_mul_left (x : E) : ⇑(isometry_equiv.mul_left x) = (*) x := rfl - -@[simp, to_additive] lemma mul_left_symm (x : E) : - (isometry_equiv.mul_left x).symm = isometry_equiv.mul_left x⁻¹ := -ext $ λ y, rfl - -variables (E) - -/-- Inversion `x ↦ x⁻¹` as an `isometry`. -/ -@[to_additive "Negation `x ↦ -x` as an `isometry`."] protected def inv : E ≃ᵢ E := -{ isometry_to_fun := isometry.of_dist_eq $ λ x y, dist_inv_inv _ _, - to_equiv := equiv.inv E } - -variables {E} - -@[simp, to_additive] lemma inv_symm : (isometry_equiv.inv E).symm = isometry_equiv.inv E := rfl -@[simp, to_additive] lemma inv_to_equiv : (isometry_equiv.inv E).to_equiv = equiv.inv E := rfl -@[simp, to_additive] lemma coe_inv : ⇑(isometry_equiv.inv E) = has_inv.inv := rfl - -end isometry_equiv - open finset @[to_additive] lemma controlled_prod_of_mem_closure {s : subgroup E} (hg : a ∈ closure (s : set E)) @@ -1201,18 +1129,6 @@ nnreal.coe_le_coe.1 $ dist_mul_mul_le a₁ a₂ b₁ b₂ lemma edist_mul_mul_le (a₁ a₂ b₁ b₂ : E) : edist (a₁ * a₂) (b₁ * b₂) ≤ edist a₁ b₁ + edist a₂ b₂ := by { simp only [edist_nndist], norm_cast, apply nndist_mul_mul_le } -@[simp, to_additive] lemma edist_mul_left (a b₁ b₂ : E) : edist (a * b₁) (a * b₂) = edist b₁ b₂ := -by simp [edist_dist] - -@[to_additive] -lemma edist_inv (a b : E) : edist a⁻¹ b = edist a b⁻¹ := by simp_rw [edist_dist, dist_inv] - -@[simp, to_additive] lemma edist_inv_inv (x y : E) : edist x⁻¹ y⁻¹ = edist x y := -by rw [edist_inv, inv_inv] - -@[simp, to_additive] lemma edist_div_left (a b₁ b₂ : E) : edist (a / b₁) (a / b₂) = edist b₁ b₂ := -by simp only [div_eq_mul_inv, edist_mul_left, edist_inv_inv] - @[to_additive] lemma nnnorm_multiset_prod_le (m : multiset E) : ‖m.prod‖₊ ≤ (m.map (λ x, ‖x‖₊)).sum := nnreal.coe_le_coe.1 $ by { push_cast, rw multiset.map_map, exact norm_multiset_prod_le _ } diff --git a/src/analysis/normed_space/basic.lean b/src/analysis/normed_space/basic.lean index 1fc2f5fff1872..adadf553ae992 100644 --- a/src/analysis/normed_space/basic.lean +++ b/src/analysis/normed_space/basic.lean @@ -80,18 +80,18 @@ lemma inv_norm_smul_mem_closed_unit_ball [normed_space ℝ β] (x : β) : by simp only [mem_closed_ball_zero_iff, norm_smul, norm_inv, norm_norm, ← div_eq_inv_mul, div_self_le_one] -lemma dist_smul [normed_space α β] (s : α) (x y : β) : dist (s • x) (s • y) = ‖s‖ * dist x y := +lemma dist_smul₀ [normed_space α β] (s : α) (x y : β) : dist (s • x) (s • y) = ‖s‖ * dist x y := by simp only [dist_eq_norm, (norm_smul _ _).symm, smul_sub] lemma nnnorm_smul [normed_space α β] (s : α) (x : β) : ‖s • x‖₊ = ‖s‖₊ * ‖x‖₊ := nnreal.eq $ norm_smul s x -lemma nndist_smul [normed_space α β] (s : α) (x y : β) : +lemma nndist_smul₀ [normed_space α β] (s : α) (x y : β) : nndist (s • x) (s • y) = ‖s‖₊ * nndist x y := -nnreal.eq $ dist_smul s x y +nnreal.eq $ dist_smul₀ s x y lemma lipschitz_with_smul [normed_space α β] (s : α) : lipschitz_with ‖s‖₊ ((•) s : β → β) := -lipschitz_with_iff_dist_le_mul.2 $ λ x y, by rw [dist_smul, coe_nnnorm] +lipschitz_with_iff_dist_le_mul.2 $ λ x y, by rw [dist_smul₀, coe_nnnorm] lemma norm_smul_of_nonneg [normed_space ℝ β] {t : ℝ} (ht : 0 ≤ t) (x : β) : ‖t • x‖ = t * ‖x‖ := by rw [norm_smul, real.norm_eq_abs, abs_of_nonneg ht] diff --git a/src/analysis/normed_space/pointwise.lean b/src/analysis/normed_space/pointwise.lean index f8981e85983f1..a236e8efe884e 100644 --- a/src/analysis/normed_space/pointwise.lean +++ b/src/analysis/normed_space/pointwise.lean @@ -29,7 +29,7 @@ begin ext y, rw mem_smul_set_iff_inv_smul_mem₀ hc, conv_lhs { rw ←inv_smul_smul₀ hc x }, - simp [← div_eq_inv_mul, div_lt_iff (norm_pos_iff.2 hc), mul_comm _ r, dist_smul], + simp [← div_eq_inv_mul, div_lt_iff (norm_pos_iff.2 hc), mul_comm _ r, dist_smul₀], end lemma smul_unit_ball {c : 𝕜} (hc : c ≠ 0) : c • ball (0 : E) (1 : ℝ) = ball (0 : E) (‖c‖) := @@ -41,7 +41,7 @@ begin ext y, rw mem_smul_set_iff_inv_smul_mem₀ hc, conv_lhs { rw ←inv_smul_smul₀ hc x }, - simp only [mem_sphere, dist_smul, norm_inv, ← div_eq_inv_mul, + simp only [mem_sphere, dist_smul₀, norm_inv, ← div_eq_inv_mul, div_eq_iff (norm_pos_iff.2 hc).ne', mul_comm r], end @@ -275,7 +275,7 @@ by rw [←cthickening_singleton _ hδ, cthickening_cthickening hε hδ, lemma ball_add_ball (hε : 0 < ε) (hδ : 0 < δ) (a b : E) : ball a ε + ball b δ = ball (a + b) (ε + δ) := -by rw [ball_add, thickening_ball hε hδ, vadd_ball, vadd_eq_add]; apply_instance +by rw [ball_add, thickening_ball hε hδ b, metric.vadd_ball, vadd_eq_add] lemma ball_sub_ball (hε : 0 < ε) (hδ : 0 < δ) (a b : E) : ball a ε - ball b δ = ball (a - b) (ε + δ) := @@ -283,7 +283,7 @@ by simp_rw [sub_eq_add_neg, neg_ball, ball_add_ball hε hδ] lemma ball_add_closed_ball (hε : 0 < ε) (hδ : 0 ≤ δ) (a b : E) : ball a ε + closed_ball b δ = ball (a + b) (ε + δ) := -by rw [ball_add, thickening_closed_ball hε hδ, vadd_ball, vadd_eq_add]; apply_instance +by rw [ball_add, thickening_closed_ball hε hδ b, metric.vadd_ball, vadd_eq_add] lemma ball_sub_closed_ball (hε : 0 < ε) (hδ : 0 ≤ δ) (a b : E) : ball a ε - closed_ball b δ = ball (a - b) (ε + δ) := @@ -291,7 +291,7 @@ by simp_rw [sub_eq_add_neg, neg_closed_ball, ball_add_closed_ball hε hδ] lemma closed_ball_add_ball (hε : 0 ≤ ε) (hδ : 0 < δ) (a b : E) : closed_ball a ε + ball b δ = ball (a + b) (ε + δ) := -by rw [add_comm, ball_add_closed_ball hδ hε, add_comm, add_comm δ]; apply_instance +by rw [add_comm, ball_add_closed_ball hδ hε b, add_comm, add_comm δ] lemma closed_ball_sub_ball (hε : 0 ≤ ε) (hδ : 0 < δ) (a b : E) : closed_ball a ε - ball b δ = ball (a - b) (ε + δ) := @@ -299,8 +299,8 @@ by simp_rw [sub_eq_add_neg, neg_ball, closed_ball_add_ball hε hδ] lemma closed_ball_add_closed_ball [proper_space E] (hε : 0 ≤ ε) (hδ : 0 ≤ δ) (a b : E) : closed_ball a ε + closed_ball b δ = closed_ball (a + b) (ε + δ) := -by rw [(is_compact_closed_ball _ _).add_closed_ball hδ, cthickening_closed_ball hδ hε, - vadd_closed_ball, vadd_eq_add, add_comm, add_comm δ]; apply_instance +by rw [(is_compact_closed_ball _ _).add_closed_ball hδ b, cthickening_closed_ball hδ hε a, + metric.vadd_closed_ball, vadd_eq_add, add_comm, add_comm δ] lemma closed_ball_sub_closed_ball [proper_space E] (hε : 0 ≤ ε) (hδ : 0 ≤ δ) (a b : E) : closed_ball a ε - closed_ball b δ = closed_ball (a - b) (ε + δ) := diff --git a/src/analysis/seminorm.lean b/src/analysis/seminorm.lean index eb76d924962bb..450906a1b3a68 100644 --- a/src/analysis/seminorm.lean +++ b/src/analysis/seminorm.lean @@ -619,7 +619,7 @@ lemma vadd_ball (p : seminorm 𝕜 E) : x +ᵥ p.ball y r = p.ball (x +ᵥ y) r := begin letI := add_group_seminorm.to_seminormed_add_comm_group p.to_add_group_seminorm, - exact vadd_ball x y r, + exact metric.vadd_ball x y r, end /-- The image of a closed ball under addition with a singleton is another closed ball. -/ @@ -627,7 +627,7 @@ lemma vadd_closed_ball (p : seminorm 𝕜 E) : x +ᵥ p.closed_ball y r = p.closed_ball (x +ᵥ y) r := begin letI := add_group_seminorm.to_seminormed_add_comm_group p.to_add_group_seminorm, - exact vadd_closed_ball x y r, + exact metric.vadd_closed_ball x y r, end end has_smul diff --git a/src/measure_theory/measure/haar_lebesgue.lean b/src/measure_theory/measure/haar_lebesgue.lean index 7cd63c3b50288..62e0626d6fb24 100644 --- a/src/measure_theory/measure/haar_lebesgue.lean +++ b/src/measure_theory/measure/haar_lebesgue.lean @@ -650,8 +650,8 @@ begin filter_upwards [self_mem_nhds_within], rintros r (rpos : 0 < r), have : closed_ball x r = {x} + r • closed_ball 0 1, - by simp only [smul_closed_ball, real.norm_of_nonneg rpos.le, zero_le_one, add_zero, mul_one, - singleton_add_closed_ball, smul_zero], + by simp only [_root_.smul_closed_ball, real.norm_of_nonneg rpos.le, zero_le_one, add_zero, + mul_one, singleton_add_closed_ball, smul_zero], simp only [this, add_haar_singleton_add_smul_div_singleton_add_smul μ rpos.ne'], simp only [add_haar_closed_ball_center, image_add_left, measure_preimage_add, singleton_add] }, have C : tendsto (λ (r : ℝ), diff --git a/src/topology/metric_space/isometric_smul.lean b/src/topology/metric_space/isometric_smul.lean new file mode 100644 index 0000000000000..f734a3402fafa --- /dev/null +++ b/src/topology/metric_space/isometric_smul.lean @@ -0,0 +1,395 @@ +/- +Copyright (c) 2022 Yury Kudryashov. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yury Kudryashov +-/ +import topology.metric_space.isometry + +/-! +# Group actions by isometries + +In this file we define two typeclasses: + +- `has_isometric_smul M X` says that `M` multiplicatively acts on a (pseudo extended) metric space + `X` by isometries; +- `has_isometric_vadd` is an additive version of `has_isometric_smul`. + +We also prove basic facts about isometric actions and define bundled isometries +`isometry_equiv.const_mul`, `isometry_equiv.mul_left`, `isometry_equiv.mul_right`, +`isometry_equiv.div_left`, `isometry_equiv.div_right`, and `isometry_equiv.inv`, as well as their +additive versions. + +If `G` is a group, then `has_isometric_smul G G` means that `G` has a left-invariant metric while +`has_isometric_smul Gᵐᵒᵖ G` means that `G` has a right-invariant metric. For a commutative group, +these two notions are equivalent. A group with a right-invariant metric can be also represented as a +`normed_group`. +-/ + +open set +open_locale ennreal pointwise + +universes u v w + +variables (M : Type u) (G : Type v) (X : Type w) + +/-- An additive action is isometric if each map `x ↦ c +ᵥ x` is an isometry. -/ +class has_isometric_vadd [pseudo_emetric_space X] [has_vadd M X] : Prop := +(isometry_vadd [] : ∀ c : M, isometry ((+ᵥ) c : X → X)) + +/-- A multiplicative action is isometric if each map `x ↦ c • x` is an isometry. -/ +@[to_additive] class has_isometric_smul [pseudo_emetric_space X] [has_smul M X] : Prop := +(isometry_smul [] : ∀ c : M, isometry ((•) c : X → X)) + +export has_isometric_vadd (isometry_vadd) has_isometric_smul (isometry_smul) + +@[priority 100, to_additive] +instance has_isometric_smul.to_has_continuous_const_smul [pseudo_emetric_space X] [has_smul M X] + [has_isometric_smul M X] : has_continuous_const_smul M X := +⟨λ c, (isometry_smul X c).continuous⟩ + +@[priority 100, to_additive] +instance has_isometric_smul.opposite_of_comm [pseudo_emetric_space X] [has_smul M X] + [has_smul Mᵐᵒᵖ X] [is_central_scalar M X] [has_isometric_smul M X] : + has_isometric_smul Mᵐᵒᵖ X := +⟨λ c x y, by simpa only [← op_smul_eq_smul] using (isometry_smul X c.unop x y)⟩ + +variables {M G X} + +section emetric + +variables [pseudo_emetric_space X] [group G] [mul_action G X] [has_isometric_smul G X] + +@[simp, to_additive] lemma edist_smul_left [has_smul M X] [has_isometric_smul M X] + (c : M) (x y : X) : + edist (c • x) (c • y) = edist x y := +isometry_smul X c x y + +@[to_additive] lemma isometry_mul_left [has_mul M] [pseudo_emetric_space M] + [has_isometric_smul M M] (a : M) : isometry ((*) a) := +isometry_smul M a + +@[simp, to_additive] lemma edist_mul_left [has_mul M] [pseudo_emetric_space M] + [has_isometric_smul M M] (a b c : M) : edist (a * b) (a * c) = edist b c := +isometry_mul_left a b c + +@[to_additive] lemma isometry_mul_right [has_mul M] [pseudo_emetric_space M] + [has_isometric_smul Mᵐᵒᵖ M] (a : M) : isometry (λ x, x * a) := +isometry_smul M (mul_opposite.op a) + +@[simp, to_additive] lemma edist_mul_right [has_mul M] [pseudo_emetric_space M] + [has_isometric_smul Mᵐᵒᵖ M] (a b c : M) : edist (a * c) (b * c) = edist a b := +isometry_mul_right c a b + +@[simp, to_additive] lemma edist_div_right [div_inv_monoid M] [pseudo_emetric_space M] + [has_isometric_smul Mᵐᵒᵖ M] (a b c : M) : edist (a / c) (b / c) = edist a b := +by simp only [div_eq_mul_inv, edist_mul_right] + +@[simp, to_additive] lemma edist_inv_inv [pseudo_emetric_space G] [has_isometric_smul G G] + [has_isometric_smul Gᵐᵒᵖ G] (a b : G) : edist a⁻¹ b⁻¹ = edist a b := +by rw [← edist_mul_left a, ← edist_mul_right _ _ b, mul_right_inv, one_mul, + inv_mul_cancel_right, edist_comm] + +@[to_additive] lemma isometry_inv [pseudo_emetric_space G] [has_isometric_smul G G] + [has_isometric_smul Gᵐᵒᵖ G] : isometry (has_inv.inv : G → G) := +edist_inv_inv + +@[to_additive] lemma edist_inv [pseudo_emetric_space G] [has_isometric_smul G G] + [has_isometric_smul Gᵐᵒᵖ G] (x y : G) : edist x⁻¹ y = edist x y⁻¹ := +by rw [← edist_inv_inv, inv_inv] + +@[simp, to_additive] lemma edist_div_left [pseudo_emetric_space G] [has_isometric_smul G G] + [has_isometric_smul Gᵐᵒᵖ G] (a b c : G) : edist (a / b) (a / c) = edist b c := +by rw [div_eq_mul_inv, div_eq_mul_inv, edist_mul_left, edist_inv_inv] + +namespace isometry_equiv + +/-- If a group `G` acts on `X` by isometries, then `isometry_equiv.const_smul` is the isometry of +`X` given by multiplication of a constant element of the group. -/ +@[to_additive "If an additive group `G` acts on `X` by isometries, then `isometry_equiv.const_vadd` +is the isometry of `X` given by addition of a constant element of the group.", simps to_equiv apply] +def const_smul (c : G) : X ≃ᵢ X := +{ to_equiv := mul_action.to_perm c, + isometry_to_fun := isometry_smul X c } + +@[simp, to_additive] +lemma const_smul_symm (c : G) : (const_smul c : X ≃ᵢ X).symm = const_smul c⁻¹ := ext $ λ _, rfl + +variables [pseudo_emetric_space G] + +/-- Multiplication `y ↦ x * y` as an `isometry_equiv`. -/ +@[to_additive "Addition `y ↦ x + y` as an `isometry_equiv`.", simps apply to_equiv] +def mul_left [has_isometric_smul G G] (c : G) : G ≃ᵢ G := +{ to_equiv := equiv.mul_left c, + isometry_to_fun := edist_mul_left c } + +@[simp, to_additive] lemma mul_left_symm [has_isometric_smul G G] (x : G) : + (mul_left x).symm = isometry_equiv.mul_left x⁻¹ := +const_smul_symm x --ext $ λ y, rfl + +/-- Multiplication `y ↦ y * x` as an `isometry_equiv`. -/ +@[to_additive "Addition `y ↦ y + x` as an `isometry_equiv`.", simps apply to_equiv] +def mul_right [has_isometric_smul Gᵐᵒᵖ G] (c : G) : G ≃ᵢ G := +{ to_equiv := equiv.mul_right c, + isometry_to_fun := λ a b, edist_mul_right a b c } + +@[simp, to_additive] lemma mul_right_symm [has_isometric_smul Gᵐᵒᵖ G] (x : G) : + (mul_right x).symm = mul_right x⁻¹ := +ext $ λ y, rfl + +/-- Division `y ↦ y / x` as an `isometry_equiv`. -/ +@[to_additive "Subtraction `y ↦ y - x` as an `isometry_equiv`.", simps apply to_equiv] +def div_right [has_isometric_smul Gᵐᵒᵖ G] (c : G) : G ≃ᵢ G := +{ to_equiv := equiv.div_right c, + isometry_to_fun := λ a b, edist_div_right a b c } + +@[simp, to_additive] lemma div_right_symm [has_isometric_smul Gᵐᵒᵖ G] (c : G) : + (div_right c).symm = mul_right c := +ext $ λ y, rfl + +variables [has_isometric_smul G G] [has_isometric_smul Gᵐᵒᵖ G] + +/-- Division `y ↦ x / y` as an `isometry_equiv`. -/ +@[to_additive "Subtraction `y ↦ x - y` as an `isometry_equiv`.", simps apply symm_apply to_equiv] +def div_left (c : G) : G ≃ᵢ G := +{ to_equiv := equiv.div_left c, + isometry_to_fun := edist_div_left c } + +variable (G) + +/-- Inversion `x ↦ x⁻¹` as an `isometry_equiv`. -/ +@[to_additive "Negation `x ↦ -x` as an `isometry_equiv`.", simps apply to_equiv] +def inv : G ≃ᵢ G := +{ to_equiv := equiv.inv G, + isometry_to_fun := edist_inv_inv } + +@[simp, to_additive] lemma inv_symm : (inv G).symm = inv G := rfl + +end isometry_equiv + +namespace emetric + +@[simp, to_additive] lemma smul_ball (c : G) (x : X) (r : ℝ≥0∞) : + c • ball x r = ball (c • x) r := +(isometry_equiv.const_smul c).image_emetric_ball _ _ + +@[simp, to_additive] lemma preimage_smul_ball (c : G) (x : X) (r : ℝ≥0∞) : + ((•) c) ⁻¹' ball x r = ball (c⁻¹ • x) r := +by rw [preimage_smul, smul_ball] + +@[simp, to_additive] lemma smul_closed_ball (c : G) (x : X) (r : ℝ≥0∞) : + c • closed_ball x r = closed_ball (c • x) r := +(isometry_equiv.const_smul c).image_emetric_closed_ball _ _ + +@[simp, to_additive] lemma preimage_smul_closed_ball (c : G) (x : X) (r : ℝ≥0∞) : + ((•) c) ⁻¹' closed_ball x r = closed_ball (c⁻¹ • x) r := +by rw [preimage_smul, smul_closed_ball] + +variables [pseudo_emetric_space G] + +@[simp, to_additive] +lemma preimage_mul_left_ball [has_isometric_smul G G] (a b : G) (r : ℝ≥0∞) : + ((*) a) ⁻¹' ball b r = ball (a⁻¹ * b) r := +preimage_smul_ball a b r + +@[simp, to_additive] +lemma preimage_mul_right_ball [has_isometric_smul Gᵐᵒᵖ G] (a b : G) (r : ℝ≥0∞) : + (λ x, x * a) ⁻¹' ball b r = ball (b / a) r := +by { rw div_eq_mul_inv, exact preimage_smul_ball (mul_opposite.op a) b r } + +@[simp, to_additive] +lemma preimage_mul_left_closed_ball [has_isometric_smul G G] (a b : G) (r : ℝ≥0∞) : + ((*) a) ⁻¹' closed_ball b r = closed_ball (a⁻¹ * b) r := +preimage_smul_closed_ball a b r + +@[simp, to_additive] +lemma preimage_mul_right_closed_ball [has_isometric_smul Gᵐᵒᵖ G] (a b : G) (r : ℝ≥0∞) : + (λ x, x * a) ⁻¹' closed_ball b r = closed_ball (b / a) r := +by { rw div_eq_mul_inv, exact preimage_smul_closed_ball (mul_opposite.op a) b r } + +end emetric + +end emetric + +@[simp, to_additive] +lemma dist_smul [pseudo_metric_space X] [has_smul M X] [has_isometric_smul M X] + (c : M) (x y : X) : dist (c • x) (c • y) = dist x y := +(isometry_smul X c).dist_eq x y + +@[simp, to_additive] +lemma nndist_smul [pseudo_metric_space X] [has_smul M X] [has_isometric_smul M X] + (c : M) (x y : X) : nndist (c • x) (c • y) = nndist x y := +(isometry_smul X c).nndist_eq x y + +@[simp, to_additive] +lemma dist_mul_left [pseudo_metric_space M] [has_mul M] [has_isometric_smul M M] + (a b c : M) : dist (a * b) (a * c) = dist b c := +dist_smul a b c + +@[simp, to_additive] +lemma nndist_mul_left [pseudo_metric_space M] [has_mul M] [has_isometric_smul M M] + (a b c : M) : nndist (a * b) (a * c) = nndist b c := +nndist_smul a b c + +@[simp, to_additive] lemma dist_mul_right [has_mul M] [pseudo_metric_space M] + [has_isometric_smul Mᵐᵒᵖ M] (a b c : M) : dist (a * c) (b * c) = dist a b := +dist_smul (mul_opposite.op c) a b + +@[simp, to_additive] +lemma nndist_mul_right [pseudo_metric_space M] [has_mul M] [has_isometric_smul Mᵐᵒᵖ M] + (a b c : M) : nndist (a * c) (b * c) = nndist a b := +nndist_smul (mul_opposite.op c) a b + +@[simp, to_additive] lemma dist_div_right [div_inv_monoid M] [pseudo_metric_space M] + [has_isometric_smul Mᵐᵒᵖ M] (a b c : M) : dist (a / c) (b / c) = dist a b := +by simp only [div_eq_mul_inv, dist_mul_right] + +@[simp, to_additive] lemma nndist_div_right [div_inv_monoid M] [pseudo_metric_space M] + [has_isometric_smul Mᵐᵒᵖ M] (a b c : M) : nndist (a / c) (b / c) = nndist a b := +by simp only [div_eq_mul_inv, nndist_mul_right] + +@[simp, to_additive] +lemma dist_inv_inv [group G] [pseudo_metric_space G] [has_isometric_smul G G] + [has_isometric_smul Gᵐᵒᵖ G] (a b : G) : dist a⁻¹ b⁻¹ = dist a b := +(isometry_equiv.inv G).dist_eq a b + +@[simp, to_additive] +lemma nndist_inv_inv [group G] [pseudo_metric_space G] [has_isometric_smul G G] + [has_isometric_smul Gᵐᵒᵖ G] (a b : G) : nndist a⁻¹ b⁻¹ = nndist a b := +(isometry_equiv.inv G).nndist_eq a b + +@[simp, to_additive] +lemma dist_div_left [group G] [pseudo_metric_space G] [has_isometric_smul G G] + [has_isometric_smul Gᵐᵒᵖ G] (a b c : G) : dist (a / b) (a / c) = dist b c := +by simp [div_eq_mul_inv] + +@[simp, to_additive] +lemma nndist_div_left [group G] [pseudo_metric_space G] [has_isometric_smul G G] + [has_isometric_smul Gᵐᵒᵖ G] (a b c : G) : nndist (a / b) (a / c) = nndist b c := +by simp [div_eq_mul_inv] + +namespace metric + +variables [pseudo_metric_space X] [group G] [mul_action G X] [has_isometric_smul G X] + +@[simp, to_additive] lemma smul_ball (c : G) (x : X) (r : ℝ) : + c • ball x r = ball (c • x) r := +(isometry_equiv.const_smul c).image_ball _ _ + +@[simp, to_additive] lemma preimage_smul_ball (c : G) (x : X) (r : ℝ) : + ((•) c) ⁻¹' ball x r = ball (c⁻¹ • x) r := +by rw [preimage_smul, smul_ball] + +@[simp, to_additive] lemma smul_closed_ball (c : G) (x : X) (r : ℝ) : + c • closed_ball x r = closed_ball (c • x) r := +(isometry_equiv.const_smul c).image_closed_ball _ _ + +@[simp, to_additive] lemma preimage_smul_closed_ball (c : G) (x : X) (r : ℝ) : + ((•) c) ⁻¹' closed_ball x r = closed_ball (c⁻¹ • x) r := +by rw [preimage_smul, smul_closed_ball] + +@[simp, to_additive] lemma smul_sphere (c : G) (x : X) (r : ℝ) : + c • sphere x r = sphere (c • x) r := +(isometry_equiv.const_smul c).image_sphere _ _ + +@[simp, to_additive] lemma preimage_smul_sphere (c : G) (x : X) (r : ℝ) : + ((•) c) ⁻¹' sphere x r = sphere (c⁻¹ • x) r := +by rw [preimage_smul, smul_sphere] + +variables [pseudo_metric_space G] + +@[simp, to_additive] +lemma preimage_mul_left_ball [has_isometric_smul G G] (a b : G) (r : ℝ) : + ((*) a) ⁻¹' ball b r = ball (a⁻¹ * b) r := +preimage_smul_ball a b r + +@[simp, to_additive] +lemma preimage_mul_right_ball [has_isometric_smul Gᵐᵒᵖ G] (a b : G) (r : ℝ) : + (λ x, x * a) ⁻¹' ball b r = ball (b / a) r := +by { rw div_eq_mul_inv, exact preimage_smul_ball (mul_opposite.op a) b r } + +@[simp, to_additive] +lemma preimage_mul_left_closed_ball [has_isometric_smul G G] (a b : G) (r : ℝ) : + ((*) a) ⁻¹' closed_ball b r = closed_ball (a⁻¹ * b) r := +preimage_smul_closed_ball a b r + +@[simp, to_additive] +lemma preimage_mul_right_closed_ball [has_isometric_smul Gᵐᵒᵖ G] (a b : G) (r : ℝ) : + (λ x, x * a) ⁻¹' closed_ball b r = closed_ball (b / a) r := +by { rw div_eq_mul_inv, exact preimage_smul_closed_ball (mul_opposite.op a) b r } + +end metric + +section instances + +variables {Y : Type*} [pseudo_emetric_space X] [pseudo_emetric_space Y] [has_smul M X] + [has_isometric_smul M X] + +@[to_additive] instance [has_smul M Y] [has_isometric_smul M Y] : + has_isometric_smul M (X × Y) := +⟨λ c, (isometry_smul X c).prod_map (isometry_smul Y c)⟩ + +@[to_additive] instance prod.has_isometric_smul' {N} + [has_mul M] [pseudo_emetric_space M] [has_isometric_smul M M] + [has_mul N] [pseudo_emetric_space N] [has_isometric_smul N N] : + has_isometric_smul (M × N) (M × N) := +⟨λ c, (isometry_smul M c.1).prod_map (isometry_smul N c.2)⟩ + +@[to_additive] instance prod.has_isometric_smul'' {N} + [has_mul M] [pseudo_emetric_space M] [has_isometric_smul Mᵐᵒᵖ M] + [has_mul N] [pseudo_emetric_space N] [has_isometric_smul Nᵐᵒᵖ N] : + has_isometric_smul (M × N)ᵐᵒᵖ (M × N) := +⟨λ c, (isometry_mul_right c.unop.1).prod_map (isometry_mul_right c.unop.2)⟩ + +@[to_additive] instance units.has_isometric_smul [monoid M] : has_isometric_smul Mˣ X := +⟨λ c, by convert isometry_smul X (c : M)⟩ + +@[to_additive] instance : has_isometric_smul M Xᵐᵒᵖ := +⟨λ c x y, by simpa only using edist_smul_left c x.unop y.unop⟩ + +@[to_additive] instance ulift.has_isometric_smul : has_isometric_smul (ulift M) X := +⟨λ c, by simpa only using isometry_smul X c.down⟩ + +@[to_additive] instance ulift.has_isometric_smul' : has_isometric_smul M (ulift X) := +⟨λ c x y, by simpa only using edist_smul_left c x.1 y.1⟩ + +@[to_additive] instance {ι} {X : ι → Type*} [fintype ι] [Π i, has_smul M (X i)] + [Π i, pseudo_emetric_space (X i)] [∀ i, has_isometric_smul M (X i)] : + has_isometric_smul M (Π i, X i) := +⟨λ c, isometry_dcomp (λ i, (•) c) (λ i, isometry_smul (X i) c)⟩ + +@[to_additive] instance pi.has_isometric_smul' {ι} {M X : ι → Type*} [fintype ι] + [Π i, has_smul (M i) (X i)] [Π i, pseudo_emetric_space (X i)] + [∀ i, has_isometric_smul (M i) (X i)] : + has_isometric_smul (Π i, M i) (Π i, X i) := +⟨λ c, isometry_dcomp (λ i, (•) (c i)) (λ i, isometry_smul _ _)⟩ + +@[to_additive] instance pi.has_isometric_smul'' {ι} {M : ι → Type*} [fintype ι] + [Π i, has_mul (M i)] [Π i, pseudo_emetric_space (M i)] [∀ i, has_isometric_smul (M i)ᵐᵒᵖ (M i)] : + has_isometric_smul (Π i, M i)ᵐᵒᵖ (Π i, M i) := +⟨λ c, isometry_dcomp (λ i (x : M i), x * c.unop i) $ λ i, isometry_mul_right _⟩ + +instance additive.has_isometric_vadd : has_isometric_vadd (additive M) X := +⟨λ c, isometry_smul X c.to_mul⟩ + +instance additive.has_isometric_vadd' [has_mul M] [pseudo_emetric_space M] + [has_isometric_smul M M] : has_isometric_vadd (additive M) (additive M) := +⟨λ c x y, edist_smul_left c.to_mul x.to_mul y.to_mul⟩ + +instance additive.has_isometric_vadd'' [has_mul M] [pseudo_emetric_space M] + [has_isometric_smul Mᵐᵒᵖ M] : has_isometric_vadd (additive M)ᵃᵒᵖ (additive M) := +⟨λ c x y, edist_smul_left (mul_opposite.op c.unop.to_mul) x.to_mul y.to_mul⟩ + +instance multiplicative.has_isometric_smul {M X} [has_vadd M X] [pseudo_emetric_space X] + [has_isometric_vadd M X]: has_isometric_smul (multiplicative M) X := +⟨λ c, isometry_vadd X c.to_add⟩ + +instance multiplicative.has_isometric_smul' [has_add M] [pseudo_emetric_space M] + [has_isometric_vadd M M] : has_isometric_smul (multiplicative M) (multiplicative M) := +⟨λ c x y, edist_vadd_left c.to_add x.to_add y.to_add⟩ + +instance multiplicative.has_isometric_vadd'' [has_add M] [pseudo_emetric_space M] + [has_isometric_vadd Mᵃᵒᵖ M] : + has_isometric_smul (multiplicative M)ᵐᵒᵖ (multiplicative M) := +⟨λ c x y, edist_vadd_left (add_opposite.op c.unop.to_add) x.to_add y.to_add⟩ + +end instances diff --git a/src/topology/metric_space/isometry.lean b/src/topology/metric_space/isometry.lean index a54e0143dbcd5..1a82febfe4a0b 100644 --- a/src/topology/metric_space/isometry.lean +++ b/src/topology/metric_space/isometry.lean @@ -76,7 +76,16 @@ lemma antilipschitz (h : isometry f) : antilipschitz_with 1 f := /-- The identity is an isometry -/ lemma _root_.isometry_id : isometry (id : α → α) := λ x y, rfl -/-- The composition of isometries is an isometry -/ +lemma prod_map {δ} [pseudo_emetric_space δ] {f : α → β} {g : γ → δ} (hf : isometry f) + (hg : isometry g) : isometry (prod.map f g) := +λ x y, by simp only [prod.edist_eq, hf.edist_eq, hg.edist_eq, prod_map] + +lemma _root_.isometry_dcomp {ι} [fintype ι] {α β : ι → Type*} [Π i, pseudo_emetric_space (α i)] + [Π i, pseudo_emetric_space (β i)] (f : Π i, α i → β i) (hf : ∀ i, isometry (f i)) : + isometry (dcomp f) := +λ x y, by simp only [edist_pi_def, (hf _).edist_eq] + +/-- The composition of isometries is an isometry. -/ theorem comp {g : β → γ} {f : α → β} (hg : isometry g) (hf : isometry f) : isometry (g ∘ f) := λ x y, (hg _ _).trans (hf _ _) From e46da4e335b8671848ac711ccb34b42538c0d800 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Wed, 1 Feb 2023 05:07:18 +0000 Subject: [PATCH 0398/1291] chore(*): add mathlib4 synchronization comments (#18342) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.big_operators.associated` * `data.W.constructions` * `data.dlist.instances` * `data.multiset.functor` * `data.sign` * `data.sum.interval` * `group_theory.coset` * `topology.algebra.order.left_right` * `topology.constructions` * `topology.continuous_on` * `topology.inseparable` * `topology.local_extr` * `topology.maps` * `topology.order` --- src/algebra/big_operators/associated.lean | 3 +++ src/data/W/constructions.lean | 3 +++ src/data/dlist/instances.lean | 3 +++ src/data/multiset/functor.lean | 3 +++ src/data/sign.lean | 3 +++ src/data/sum/interval.lean | 3 +++ src/group_theory/coset.lean | 3 +++ src/topology/algebra/order/left_right.lean | 3 +++ src/topology/constructions.lean | 3 +++ src/topology/continuous_on.lean | 3 +++ src/topology/inseparable.lean | 3 +++ src/topology/local_extr.lean | 3 +++ src/topology/maps.lean | 3 +++ src/topology/order.lean | 3 +++ 14 files changed, 42 insertions(+) diff --git a/src/algebra/big_operators/associated.lean b/src/algebra/big_operators/associated.lean index 8412711ed5ac5..49c80173ad268 100644 --- a/src/algebra/big_operators/associated.lean +++ b/src/algebra/big_operators/associated.lean @@ -10,6 +10,9 @@ import algebra.big_operators.finsupp /-! # Products of associated, prime, and irreducible elements. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains some theorems relating definitions in `algebra.associated` and products of multisets, finsets, and finsupps. diff --git a/src/data/W/constructions.lean b/src/data/W/constructions.lean index 654dfd030b995..3e1698d34aa5e 100644 --- a/src/data/W/constructions.lean +++ b/src/data/W/constructions.lean @@ -8,6 +8,9 @@ import data.W.basic /-! # Examples of W-types +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We take the view of W types as inductive types. Given `α : Type` and `β : α → Type`, the W type determined by this data, `W_type β`, is the inductively with constructors from `α` and arities of each constructor `a : α` given by `β a`. diff --git a/src/data/dlist/instances.lean b/src/data/dlist/instances.lean index a1ae395cf3291..62d1d97465dfa 100644 --- a/src/data/dlist/instances.lean +++ b/src/data/dlist/instances.lean @@ -9,6 +9,9 @@ import control.traversable.instances /-! # Traversable instance for dlists +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file provides the equivalence between `list α` and `dlist α` and the traversable instance for `dlist`. -/ diff --git a/src/data/multiset/functor.lean b/src/data/multiset/functor.lean index 835d378df932c..eac8ae74d286b 100644 --- a/src/data/multiset/functor.lean +++ b/src/data/multiset/functor.lean @@ -9,6 +9,9 @@ import control.traversable.instances /-! # Functoriality of `multiset`. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ universes u diff --git a/src/data/sign.lean b/src/data/sign.lean index 0e2013bcf4553..9420ce687e910 100644 --- a/src/data/sign.lean +++ b/src/data/sign.lean @@ -11,6 +11,9 @@ import tactic.derive_fintype /-! # Sign function +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the sign function for types with zero and a decidable less-than relation, and proves some basic theorems about it. -/ diff --git a/src/data/sum/interval.lean b/src/data/sum/interval.lean index 77e525f926504..5bd433861ad78 100644 --- a/src/data/sum/interval.lean +++ b/src/data/sum/interval.lean @@ -9,6 +9,9 @@ import order.locally_finite /-! # Finite intervals in a disjoint union +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file provides the `locally_finite_order` instance for the disjoint sum of two orders. ## TODO diff --git a/src/group_theory/coset.lean b/src/group_theory/coset.lean index 69843a35df0c7..01fc5e6c7e3e5 100644 --- a/src/group_theory/coset.lean +++ b/src/group_theory/coset.lean @@ -13,6 +13,9 @@ import tactic.group /-! # Cosets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file develops the basic theory of left and right cosets. ## Main definitions diff --git a/src/topology/algebra/order/left_right.lean b/src/topology/algebra/order/left_right.lean index fe98ae82544fe..2b43409de367c 100644 --- a/src/topology/algebra/order/left_right.lean +++ b/src/topology/algebra/order/left_right.lean @@ -8,6 +8,9 @@ import topology.continuous_on /-! # Left and right continuity +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove a few lemmas about left and right continuous functions: * `continuous_within_at_Ioi_iff_Ici`: two definitions of right continuity diff --git a/src/topology/constructions.lean b/src/topology/constructions.lean index 6e64a6c0ba3b8..b59b18db550df 100644 --- a/src/topology/constructions.lean +++ b/src/topology/constructions.lean @@ -10,6 +10,9 @@ import order.filter.pi /-! # Constructions of new topological spaces from old ones +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file constructs products, sums, subtypes and quotients of topological spaces and sets up their basic theory, such as criteria for maps into or out of these constructions to be continuous; descriptions of the open sets, neighborhood filters, diff --git a/src/topology/continuous_on.lean b/src/topology/continuous_on.lean index f2a69334ad418..e9e6b09217110 100644 --- a/src/topology/continuous_on.lean +++ b/src/topology/continuous_on.lean @@ -8,6 +8,9 @@ import topology.constructions /-! # Neighborhoods and continuity relative to a subset +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines relative versions * `nhds_within` of `nhds` diff --git a/src/topology/inseparable.lean b/src/topology/inseparable.lean index f5dd2759c0953..0d8200d603bbf 100644 --- a/src/topology/inseparable.lean +++ b/src/topology/inseparable.lean @@ -10,6 +10,9 @@ import tactic.tfae /-! # Inseparable points in a topological space +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define * `specializes` (notation: `x ⤳ y`) : a relation saying that `𝓝 x ≤ 𝓝 y`; diff --git a/src/topology/local_extr.lean b/src/topology/local_extr.lean index 9ed89b5e3a917..4abf12035b58c 100644 --- a/src/topology/local_extr.lean +++ b/src/topology/local_extr.lean @@ -9,6 +9,9 @@ import topology.continuous_on /-! # Local extrema of functions on topological spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions This file defines special versions of `is_*_filter f a l`, `*=min/max/extr`, diff --git a/src/topology/maps.lean b/src/topology/maps.lean index 3b3cc60f64cfb..1f6641fd51246 100644 --- a/src/topology/maps.lean +++ b/src/topology/maps.lean @@ -9,6 +9,9 @@ import topology.nhds_set /-! # Specific classes of maps between topological spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file introduces the following properties of a map `f : X → Y` between topological spaces: * `is_open_map f` means the image of an open set under `f` is open. diff --git a/src/topology/order.lean b/src/topology/order.lean index 04914a90301e1..23445685a8349 100644 --- a/src/topology/order.lean +++ b/src/topology/order.lean @@ -8,6 +8,9 @@ import topology.tactic /-! # Ordering on topologies and (co)induced topologies +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Topologies on a fixed type `α` are ordered, by reverse inclusion. That is, for topologies `t₁` and `t₂` on `α`, we write `t₁ ≤ t₂` if every set open in `t₂` is also open in `t₁`. From 78ac1db34d90519b8e8f7bfe39cb264a6a117db2 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Wed, 1 Feb 2023 07:41:09 +0000 Subject: [PATCH 0399/1291] feat(number_theory/number_field/embeddings): add infinite_places.eq_iff (#17285) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This is the proof that two embeddings of a field into `ℂ` define the same (infinite) place iff they are equal or complex conjugate. This PR also contains definitions of complex and real infinite places and some basic results about these. Two lemmas about subfields of `ℂ` are in the new file `topology/instances/complex.lean` but they might belong elsewhere. --- .../number_field/embeddings.lean | 84 ++++++++++++++++ src/topology/instances/complex.lean | 97 +++++++++++++++++++ 2 files changed, 181 insertions(+) create mode 100644 src/topology/instances/complex.lean diff --git a/src/number_theory/number_field/embeddings.lean b/src/number_theory/number_field/embeddings.lean index def8c20de89d5..0abf10cf00fe2 100644 --- a/src/number_theory/number_field/embeddings.lean +++ b/src/number_theory/number_field/embeddings.lean @@ -8,6 +8,7 @@ import analysis.complex.polynomial import data.complex.basic import field_theory.minpoly.is_integrally_closed import number_theory.number_field.basic +import topology.instances.complex /-! # Embeddings of number fields @@ -245,6 +246,89 @@ subtype.ext (w.2).some_spec lemma pos_iff (w : infinite_place K) (x : K) : 0 < w x ↔ x ≠ 0 := absolute_value.pos_iff w.1 +@[simp] +lemma mk_conjugate_eq (φ : K →+* ℂ) : + mk (complex_embedding.conjugate φ) = mk φ := +begin + ext x, + exact congr_fun (congr_arg coe_fn (complex_embedding.place_conjugate φ)) x, +end + +@[simp] +lemma mk_eq_iff {φ ψ : K →+* ℂ} : + mk φ = mk ψ ↔ φ = ψ ∨ complex_embedding.conjugate φ = ψ := +begin + split, + { -- We prove that the map ψ ∘ φ⁻¹ between φ(K) and ℂ is uniform continuous, thus it is either the + -- inclusion or the complex conjugation using complex.uniform_continuous_ring_hom_eq_id_or_conj + intro h₀, + obtain ⟨j, hiφ⟩ := φ.injective.has_left_inverse, + let ι := ring_equiv.of_left_inverse hiφ, + have hlip : lipschitz_with 1 (ring_hom.comp ψ ι.symm.to_ring_hom), + { change lipschitz_with 1 (ψ ∘ ι.symm), + apply lipschitz_with.of_dist_le_mul, + intros x y, + rw [nonneg.coe_one, one_mul, normed_field.dist_eq, ← map_sub, ← map_sub], + apply le_of_eq, + suffices : ‖φ ((ι.symm) (x - y))‖ = ‖ψ ((ι.symm) (x - y))‖, + { rw [← this, ← ring_equiv.of_left_inverse_apply hiφ _ , ring_equiv.apply_symm_apply ι _], + refl, }, + exact congr_fun (congr_arg coe_fn h₀) _, }, + cases (complex.uniform_continuous_ring_hom_eq_id_or_conj φ.field_range hlip.uniform_continuous), + { left, ext1 x, + convert (congr_fun h (ι x)).symm, + exact (ring_equiv.apply_symm_apply ι.symm x).symm, }, + { right, ext1 x, + convert (congr_fun h (ι x)).symm, + exact (ring_equiv.apply_symm_apply ι.symm x).symm, }}, + { rintros (⟨h⟩ | ⟨h⟩), + { exact congr_arg mk h, }, + { rw ← mk_conjugate_eq, + exact congr_arg mk h, }}, +end + +/-- An infinite place is real if it is defined by a real embedding. -/ +def is_real (w : infinite_place K) : Prop := + ∃ φ : K →+* ℂ, complex_embedding.is_real φ ∧ mk φ = w + +/-- An infinite place is complex if it is defined by a complex (ie. not real) embedding. -/ +def is_complex (w : infinite_place K) : Prop := + ∃ φ : K →+* ℂ, ¬ complex_embedding.is_real φ ∧ mk φ = w + +lemma _root_.number_field.complex_embeddings.is_real.embedding_mk {φ : K →+* ℂ} + (h : complex_embedding.is_real φ) : + embedding (mk φ) = φ := +begin + have := mk_eq_iff.mp (mk_embedding (mk φ)).symm, + rwa [complex_embedding.is_real_iff.mp h, or_self, eq_comm] at this, +end + +lemma is_real_iff {w : infinite_place K} : + is_real w ↔ complex_embedding.is_real (embedding w) := +begin + split, + { rintros ⟨φ, ⟨hφ, rfl⟩⟩, + rwa _root_.number_field.complex_embeddings.is_real.embedding_mk hφ, }, + { exact λ h, ⟨embedding w, h, mk_embedding w⟩, }, +end + +lemma is_complex_iff {w : infinite_place K} : + is_complex w ↔ ¬ complex_embedding.is_real (embedding w) := +begin + split, + { rintros ⟨φ, ⟨hφ, rfl⟩⟩, + contrapose! hφ, + cases mk_eq_iff.mp (mk_embedding (mk φ)), + { rwa ← h, }, + { rw ← complex_embedding.is_real_conjugate_iff at hφ, + rwa ← h, }}, + { exact λ h, ⟨embedding w, h, mk_embedding w⟩, }, + end + +lemma not_is_real_iff_is_complex {w : infinite_place K} : + ¬ is_real w ↔ is_complex w := +by rw [is_complex_iff, is_real_iff] + end number_field.infinite_place end infinite_place diff --git a/src/topology/instances/complex.lean b/src/topology/instances/complex.lean new file mode 100644 index 0000000000000..5b88f2cf55212 --- /dev/null +++ b/src/topology/instances/complex.lean @@ -0,0 +1,97 @@ +/- +Copyright (c) 2022 Xavier Roblot. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Xavier Roblot +-/ + +import topology.algebra.uniform_field +import analysis.complex.basic +import field_theory.adjoin + +/-! +# Some results about the topology of ℂ +-/ + +section complex_subfield + +open complex set + +open_locale complex_conjugate + +/-- The only closed subfields of `ℂ` are `ℝ` and `ℂ`. -/ +lemma complex.subfield_eq_of_closed {K : subfield ℂ} (hc : is_closed (K : set ℂ)) : + K = of_real.field_range ∨ K = ⊤ := +begin + suffices : range (coe : ℝ → ℂ) ⊆ K, + { rw [range_subset_iff, ← coe_algebra_map] at this, + have := (subalgebra.is_simple_order_of_finrank finrank_real_complex).eq_bot_or_eq_top + (subfield.to_intermediate_field K this).to_subalgebra, + simp_rw ← set_like.coe_set_eq at this ⊢, + convert this using 2, + simpa only [ring_hom.coe_field_range, algebra.coe_bot, coe_algebra_map], }, + suffices : range (coe : ℝ → ℂ) ⊆ closure (set.range ((coe : ℝ → ℂ) ∘ (coe : ℚ → ℝ))), + { refine subset_trans this _, + rw ← is_closed.closure_eq hc, + apply closure_mono, + rintros _ ⟨_, rfl⟩, + simp only [function.comp_app, of_real_rat_cast, set_like.mem_coe, subfield_class.coe_rat_mem] }, + nth_rewrite 1 range_comp, + refine subset_trans _ (image_closure_subset_closure_image continuous_of_real), + rw dense_range.closure_range rat.dense_embedding_coe_real.dense, + simp only [image_univ], +end + +/-- Let `K` a subfield of `ℂ` and let `ψ : K →+* ℂ` a ring homomorphism. Assume that `ψ` is uniform +continuous, then `ψ` is either the inclusion map or the composition of the inclusion map with the +complex conjugation. -/ +lemma complex.uniform_continuous_ring_hom_eq_id_or_conj (K : subfield ℂ) {ψ : K →+* ℂ} + (hc : uniform_continuous ψ) : ψ.to_fun = K.subtype ∨ ψ.to_fun = conj ∘ K.subtype := +begin + letI : topological_division_ring ℂ := topological_division_ring.mk, + letI : topological_ring K.topological_closure := + subring.topological_ring K.topological_closure.to_subring, + set ι : K → K.topological_closure := subfield.inclusion K.le_topological_closure, + have ui : uniform_inducing ι := + ⟨ by { erw [uniformity_subtype, uniformity_subtype, filter.comap_comap], congr, } ⟩, + let di := ui.dense_inducing _, + { -- extψ : closure(K) →+* ℂ is the extension of ψ : K →+* ℂ + let extψ := dense_inducing.extend_ring_hom ui di.dense hc, + haveI := (uniform_continuous_uniformly_extend ui di.dense hc).continuous, + cases complex.subfield_eq_of_closed (subfield.is_closed_topological_closure K), + { left, + let j := ring_equiv.subfield_congr h, + -- ψ₁ is the continuous ring hom `ℝ →+* ℂ` constructed from `j : closure (K) ≃+* ℝ` + -- and `extψ : closure (K) →+* ℂ` + let ψ₁ := ring_hom.comp extψ (ring_hom.comp j.symm.to_ring_hom of_real.range_restrict), + ext1 x, + rsuffices ⟨r, hr⟩ : ∃ r : ℝ, of_real.range_restrict r = j (ι x), + { have := ring_hom.congr_fun + (ring_hom_eq_of_real_of_continuous (by continuity! : continuous ψ₁)) r, + rw [ring_hom.comp_apply, ring_hom.comp_apply, hr, ring_equiv.to_ring_hom_eq_coe] at this, + convert this using 1, + { exact (dense_inducing.extend_eq di hc.continuous _).symm, }, + { rw [← of_real.coe_range_restrict, hr], refl, }}, + obtain ⟨r, hr⟩ := set_like.coe_mem (j (ι x)), + exact ⟨r, subtype.ext hr⟩, }, + { -- ψ₁ is the continuous ring hom `ℂ →+* ℂ` constructed from `closure (K) ≃+* ℂ` + -- and `extψ : closure (K) →+* ℂ` + let ψ₁ := ring_hom.comp extψ (ring_hom.comp (ring_equiv.subfield_congr h).symm.to_ring_hom + (@subfield.top_equiv ℂ _).symm.to_ring_hom), + cases ring_hom_eq_id_or_conj_of_continuous (by continuity! : continuous ψ₁) with h h, + { left, ext1 z, + convert (ring_hom.congr_fun h z) using 1, + exact (dense_inducing.extend_eq di hc.continuous z).symm, }, + { right, ext1 z, + convert (ring_hom.congr_fun h z) using 1, + exact (dense_inducing.extend_eq di hc.continuous z).symm, }}}, + { let j : { x // x ∈ closure (id '' {x | (K : set ℂ) x })} → (K.topological_closure : set ℂ) := + λ x, ⟨x, by { convert x.prop, simpa only [id.def, set.image_id'], }⟩, + convert dense_range.comp (function.surjective.dense_range _) + (dense_embedding.subtype (dense_embedding_id) (K : set ℂ)).dense + (by continuity : continuous j), + rintros ⟨y, hy⟩, + use ⟨y, by { convert hy, simpa only [id.def, set.image_id'], }⟩, + simp only [subtype.mk_eq_mk, subtype.coe_mk], } +end + +end complex_subfield From 980755c33b9168bc82f774f665eaa27878140fac Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Wed, 1 Feb 2023 07:41:10 +0000 Subject: [PATCH 0400/1291] feat(analysis/special_functions/trigonometric): Euler's infinite product for sine (part b) (#18281) Co-authored-by: sgouezel --- .../trigonometric/euler_sine_prod.lean | 68 +++++++++- src/data/real/pi/wallis.lean | 127 +----------------- 2 files changed, 69 insertions(+), 126 deletions(-) diff --git a/src/analysis/special_functions/trigonometric/euler_sine_prod.lean b/src/analysis/special_functions/trigonometric/euler_sine_prod.lean index 21fb9a9b5eb98..902152adc5abf 100644 --- a/src/analysis/special_functions/trigonometric/euler_sine_prod.lean +++ b/src/analysis/special_functions/trigonometric/euler_sine_prod.lean @@ -4,9 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: David Loeffler -/ -import analysis.special_functions.integrals -import analysis.special_functions.trigonometric.bounds import data.real.pi.wallis +import measure_theory.integral.peak_function /-! # Euler's infinite product for the sine function @@ -252,4 +251,69 @@ end end integral_recursion + +/-! ## Conclusion of the proof + +The main theorem `complex.tendsto_euler_sin_prod`, and its real variant +`real.tendsto_euler_sin_prod`, now follow by combining `sin_pi_mul_eq` with a lemma +stating that the sequence of measures on `[0, π/2]` given by integration against `cos x ^ n` +(suitably normalised) tends to the Dirac measure at 0, as a special case of the general result +`tendsto_set_integral_pow_smul_of_unique_maximum_of_is_compact_of_continuous_on`. -/ + +lemma tendsto_integral_cos_pow_mul_div {f : ℝ → ℂ} (hf : continuous_on f (Icc 0 (π/2))) : + tendsto (λ (n : ℕ), (∫ x:ℝ in 0..π/2, ↑(cos x) ^ n * f x) / ↑(∫ x:ℝ in 0..π/2, (cos x) ^ n)) + at_top (𝓝 $ f 0) := +begin + simp_rw [div_eq_inv_mul _ (coe _), ←complex.of_real_inv, integral_of_le (pi_div_two_pos.le), + ←measure_theory.integral_Icc_eq_integral_Ioc, ←complex.of_real_pow, ←complex.real_smul], + have c_lt : ∀ (y : ℝ), y ∈ Icc 0 (π / 2) → y ≠ 0 → cos y < cos 0, from λ y hy hy', + cos_lt_cos_of_nonneg_of_le_pi_div_two (le_refl 0) hy.2 (lt_of_le_of_ne hy.1 hy'.symm), + have c_nonneg : ∀ (x : ℝ), x ∈ Icc 0 (π / 2) → 0 ≤ cos x, from λ x hx, cos_nonneg_of_mem_Icc + ((Icc_subset_Icc_left (neg_nonpos_of_nonneg pi_div_two_pos.le)) hx), + have c_zero_pos : 0 < cos 0, by { rw cos_zero, exact zero_lt_one }, + have zero_mem : (0:ℝ) ∈ closure (interior (Icc 0 (π / 2))), + { rw [interior_Icc, closure_Ioo pi_div_two_pos.ne, left_mem_Icc], + exact pi_div_two_pos.le }, + exact tendsto_set_integral_pow_smul_of_unique_maximum_of_is_compact_of_continuous_on + is_compact_Icc continuous_on_cos c_lt c_nonneg c_zero_pos zero_mem hf +end + +/-- Euler's infinite product formula for the complex sine function. -/ +lemma _root_.complex.tendsto_euler_sin_prod (z : ℂ) : + tendsto (λ n:ℕ, ↑π * z * (∏ j in finset.range n, (1 - z ^ 2 / (j + 1) ^ 2))) + at_top (𝓝 $ complex.sin (π * z)) := +begin + have A : tendsto (λ n:ℕ, ↑π * z * (∏ j in finset.range n, (1 - z ^ 2 / (j + 1) ^ 2)) * + (∫ x in 0..π / 2, complex.cos (2 * z * x) * cos x ^ (2 * n)) / + ↑∫ x in 0..π / 2, cos x ^ (2 * n)) + at_top (𝓝 $ _) := tendsto.congr (λ n, (sin_pi_mul_eq z n)) tendsto_const_nhds, + have : 𝓝 (complex.sin (π * z)) = 𝓝 (complex.sin (π * z) * 1) := by rw mul_one, + simp_rw [this, mul_div_assoc] at A, + convert (tendsto_mul_iff_of_ne_zero _ one_ne_zero).mp A, + suffices : tendsto (λ n:ℕ, (∫ x:ℝ in 0..π/2, complex.cos (2 * z * x) * cos x ^ n) + / ↑(∫ x:ℝ in 0..π/2, cos x ^ n)) at_top (𝓝 1), + from this.comp (tendsto_id.const_mul_at_top' zero_lt_two), + have : continuous_on (λ x:ℝ, complex.cos (2 * z * x)) (Icc 0 (π/2)), from + (complex.continuous_cos.comp (continuous_const.mul complex.continuous_of_real)).continuous_on, + convert tendsto_integral_cos_pow_mul_div this, + { ext1 n, congr' 2 with x:1, rw mul_comm }, + { rw [complex.of_real_zero, mul_zero, complex.cos_zero] }, +end + +/-- Euler's infinite product formula for the real sine function. -/ +lemma _root_.real.tendsto_euler_sin_prod (x : ℝ) : + tendsto (λ n:ℕ, π * x * (∏ j in finset.range n, (1 - x ^ 2 / (j + 1) ^ 2))) + at_top (𝓝 $ sin (π * x)) := +begin + convert (complex.continuous_re.tendsto _).comp (complex.tendsto_euler_sin_prod x), + { ext1 n, + rw [function.comp_app, ←complex.of_real_mul, complex.of_real_mul_re], + suffices : ∏ (j : ℕ) in finset.range n, (1 - (x:ℂ) ^ 2 / (↑j + 1) ^ 2) = + ↑∏ (j : ℕ) in finset.range n, (1 - x ^ 2 / (↑j + 1) ^ 2), by rw [this, complex.of_real_re], + rw complex.of_real_prod, + refine finset.prod_congr (by refl) (λ n hn, _), + norm_cast }, + { rw [←complex.of_real_mul, ←complex.of_real_sin, complex.of_real_re] } +end + end euler_sine diff --git a/src/data/real/pi/wallis.lean b/src/data/real/pi/wallis.lean index c58e97b4aad7c..2f114d5606c4e 100644 --- a/src/data/real/pi/wallis.lean +++ b/src/data/real/pi/wallis.lean @@ -26,9 +26,6 @@ algebraic manipulation. * `real.wallis.W`: the product of the first `k` terms in Wallis' formula for `π`. * `real.wallis.W_eq_integral_sin_pow_div_integral_sin_pow`: express `W n` as a ratio of integrals. * `real.wallis.W_le` and `real.wallis.le_W`: upper and lower bounds for `W n`. -* `real.wallis.integral_sin_pow_odd_sq_eq` and `real.wallis.integral_sin_pow_even_sq_eq`: formulas - for `(∫ x in 0..π, sin x ^ n) ^ 2` in terms of `W`. -* `integral_sin_pow_le` and `le_integral_sin_pow`: bounds for `∫ x in 0..π, sin x ^ n`. * `real.tendsto_prod_pi_div_two`: the Wallis product formula. -/ @@ -92,7 +89,7 @@ begin rw [←le_div_iff pi_div_two_pos, div_eq_inv_mul (W k) _], rw [W_eq_integral_sin_pow_div_integral_sin_pow, le_div_iff (integral_sin_pow_pos _)], convert integral_sin_pow_succ_le (2 * k + 1), - rw [integral_sin_pow (2 * k)], + rw integral_sin_pow (2 * k), simp only [sin_zero, zero_pow', ne.def, nat.succ_ne_zero, not_false_iff, zero_mul, sin_pi, tsub_zero, nat.cast_mul, nat.cast_bit0, algebra_map.coe_one, zero_div, zero_add], end @@ -110,134 +107,16 @@ begin simp_rw h, refine (tendsto_const_nhds.div_at_top _).const_sub _, refine tendsto.at_top_add _ tendsto_const_nhds, - exact tendsto_coe_nat_at_top_at_top.const_mul_at_top two_pos, -end - -lemma W_eq_mul_sq (k : ℕ) : - W k = (2 * k + 1) * (∏ i in range k, ((2:ℝ) * i + 2) / (2 * i + 3)) ^ 2 := -begin - induction k with k hk, - { simp [W], }, - { unfold W at *, - rw [prod_range_succ, prod_range_succ, hk], - suffices : ∀ (x : ℝ), - (2 * ↑k + 1) * x ^ 2 * ((2 * ↑k + 2) / (2 * ↑k + 1) * ((2 * ↑k + 2) / (2 * ↑k + 3))) = - (2 * ↑(k.succ) + 1) * (x * ((2 * ↑k + 2) / (2 * ↑k + 3))) ^ 2, - { rw this }, - intro x, - have a : (2 * ↑k + 1 : ℝ) ≠ 0, by positivity, - have b : (2 * ↑k + 3 : ℝ) ≠ 0, by positivity, - field_simp, ring } -end - -lemma integral_sin_pow_odd_sq_eq (k : ℕ) : - (∫ x in 0..π, sin x ^ (2 * k + 1)) ^ 2 = 4 * W k / (2 * k + 1) := -begin - rw integral_sin_pow_odd, - have B := W_eq_mul_sq k, - rw [mul_comm (2 * (k:ℝ) + 1) _, ←div_eq_iff] at B, - { rw [mul_pow, ←B], - ring }, - { positivity }, -end - -lemma integral_sin_pow_even_sq_eq (k : ℕ) : - (∫ x in 0..π, sin x ^ (2 * k)) ^ 2 = π ^ 2 / (2 * k + 1) / W k := -begin - induction k with k hk, - { dsimp only [W], - simp }, - { have np : 0 < 2 * (k:ℝ) + 1, by positivity, - rw [nat.succ_eq_add_one, mul_add 2 k 1, mul_one, integral_sin_pow, sin_zero, sin_pi, - zero_pow (nat.add_pos_right _ zero_lt_one), zero_mul, zero_mul, sub_zero, zero_div, zero_add, - mul_pow, hk, W_succ, nat.cast_add_one, nat.cast_mul, mul_add, mul_one, - add_assoc (2 * (k:ℝ)) 2 1, (by ring : (2:ℝ) + 1 = 3), sq], - have np2 : 2 * (k:ℝ) + 2 ≠ 0, by positivity, - have np3 : 2 * (k:ℝ) + 3 ≠ 0, by positivity, - field_simp [np.ne', (W_pos k).ne'], - ring } + exact tendsto_coe_nat_at_top_at_top.const_mul_at_top two_pos end end wallis end real -open real real.wallis - -section integral_sin_pow_bounds -/-! ## Bounds for integrals of `sin x ^ n` - -Explicit `O(1/√n)` bounds for `∫ x in 0..π, sin x ^ n`, as a by-product of the proof of Wallis' -formula for `π`. -/ - -lemma integral_sin_pow_odd_le (n : ℕ) : - ∫ x in 0..π, sin x ^ (2 * n + 1) ≤ sqrt (2 * π / (2 * n + 1)) := -begin - have np : 0 < 2 * (n:ℝ) + 1, by positivity, - rw [le_sqrt (integral_sin_pow_pos _).le (div_pos two_pi_pos np).le, integral_sin_pow_odd_sq_eq], - apply div_le_div_of_le np.le, - rw ←le_div_iff' (by norm_num : 0 < (4:ℝ)), - convert W_le n using 1, - ring, -end - -lemma integral_sin_pow_even_le (n : ℕ) : - ∫ x in 0..π, sin x ^ (2 * n) ≤ sqrt (2 * π * (2 * n + 2) / (2 * n + 1) ^ 2) := -begin - have np : 0 < 2 * (n:ℝ) + 1, by positivity, - have np' : 0 < 2 * (n:ℝ) + 2, by positivity, - rw le_sqrt (integral_sin_pow_pos _).le, - swap, { refine div_nonneg _ (sq_nonneg _), exact mul_nonneg (two_pi_pos).le np'.le }, - rw [integral_sin_pow_even_sq_eq, div_le_iff (W_pos n), ←div_le_iff'], - swap, { refine div_pos _ (sq_pos_of_pos np), exact mul_pos two_pi_pos np' }, - convert le_W n, - field_simp [np.ne', np'.ne', pi_pos.ne'], - ring, -end - -lemma integral_sin_pow_le {n : ℕ} (hn : n ≠ 0) : ∫ x in 0..π, sin x ^ n ≤ sqrt (2 * π / n) := -begin - -- this is a slightly weaker bound than `integral_sin_pow_even_le` for even `n`, but uniform in - -- its statement - obtain ⟨k, hk⟩ := nat.even_or_odd' n, - rcases hk with rfl | rfl, - { refine le_trans (integral_sin_pow_even_le k) _, - apply sqrt_le_sqrt, - rw [div_le_div_iff, mul_assoc, mul_le_mul_left two_pi_pos], - rotate, { positivity }, { positivity }, - have : (2 * (k:ℝ) + 2) * ((2 * k : ℕ) : ℝ) = (2 * k + 1) ^ 2 - 1, - { push_cast, ring, }, - rw [this, sub_le_self_iff], - exact zero_le_one }, - { convert integral_sin_pow_odd_le k using 3, - rw [nat.cast_add, nat.cast_mul, nat.cast_two, nat.cast_one] }, -end - -lemma le_integral_sin_pow (n : ℕ) : sqrt (2 * π / (n + 1)) ≤ ∫ x in 0..π, sin x ^ n := -begin - refine sqrt_le_iff.mpr ⟨(integral_sin_pow_pos _).le, _⟩, - obtain ⟨k, hk⟩ := nat.even_or_odd' n, - have np : 0 < 2 * (k:ℝ) + 1, by positivity, - have np' : 2 * (k:ℝ) + 2 ≠ 0, by positivity, - rcases hk with rfl | rfl, - { rw [integral_sin_pow_even_sq_eq, le_div_iff (W_pos _), nat.cast_mul, nat.cast_two, - ←le_div_iff' (div_pos two_pi_pos np)], - convert W_le k using 1, - field_simp [np.ne', np', pi_pos.ne'], - ring }, - { rw [nat.cast_add, nat.cast_mul, nat.cast_two, nat.cast_one, - (by ring : (2:ℝ) * k + 1 + 1 = 2 * k + 2), integral_sin_pow_odd_sq_eq, le_div_iff np, - ←div_le_iff' (by positivity : 0 < (4:ℝ))], - convert le_W k, - field_simp [np.ne', np'], - ring }, -end - -end integral_sin_pow_bounds - /-- Wallis' product formula for `π / 2`. -/ theorem real.tendsto_prod_pi_div_two : tendsto (λ k, ∏ i in range k, (((2:ℝ) * i + 2) / (2 * i + 1)) * ((2 * i + 2) / (2 * i + 3))) at_top (𝓝 (π/2)) := -tendsto_W_nhds_pi_div_two +real.wallis.tendsto_W_nhds_pi_div_two From 2f4cdce0c2f2f3b8cd58f05d556d03b468e1eb2e Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 1 Feb 2023 09:29:16 +0000 Subject: [PATCH 0401/1291] feat(linear_algebra/affine_space/basis): `reindex` API (#18190) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Rename and change arguments to `affine_basis.comp_equiv` so that it matches `basis.reindex`. Provide lemmas for its interaction with other `affine_basis` constructions. Make `affine_basis` follow the `fun_like` design, replacing `affine_basis.points` by a function coercion. Fix a few names all around: * `affine_basis.coord_apply_neq` → `affine_basis.coord_apply_ne` * `convex_hull_affine_basis_eq_nonneg_barycentric` → `affine_basis.convex_hull_eq_nonneg_coord` --- src/analysis/convex/combination.lean | 9 +- src/analysis/inner_product_space/pi_L2.lean | 4 +- .../normed_space/add_torsor_bases.lean | 12 +- src/data/set/image.lean | 32 +++-- src/linear_algebra/affine_space/basis.lean | 115 +++++++++++------- .../affine_space/finite_dimensional.lean | 8 +- src/linear_algebra/affine_space/matrix.lean | 24 ++-- src/linear_algebra/basis.lean | 51 ++++---- src/linear_algebra/matrix/basis.lean | 8 +- src/linear_algebra/matrix/spectrum.lean | 4 +- 10 files changed, 155 insertions(+), 112 deletions(-) diff --git a/src/analysis/convex/combination.lean b/src/analysis/convex/combination.lean index 8aac405b7832b..f21d458d49b99 100644 --- a/src/analysis/convex/combination.lean +++ b/src/analysis/convex/combination.lean @@ -461,19 +461,18 @@ lemma mem_Icc_of_mem_std_simplex (hf : f ∈ std_simplex R ι) (x) : /-- The convex hull of an affine basis is the intersection of the half-spaces defined by the corresponding barycentric coordinates. -/ -lemma convex_hull_affine_basis_eq_nonneg_barycentric {ι : Type*} (b : affine_basis ι R E) : - convex_hull R (range b.points) = { x | ∀ i, 0 ≤ b.coord i x } := +lemma affine_basis.convex_hull_eq_nonneg_coord {ι : Type*} (b : affine_basis ι R E) : + convex_hull R (range b) = {x | ∀ i, 0 ≤ b.coord i x} := begin rw convex_hull_range_eq_exists_affine_combination, ext x, - split, + refine ⟨_, λ hx, _⟩, { rintros ⟨s, w, hw₀, hw₁, rfl⟩ i, by_cases hi : i ∈ s, { rw b.coord_apply_combination_of_mem hi hw₁, exact hw₀ i hi, }, { rw b.coord_apply_combination_of_not_mem hi hw₁, }, }, - { intros hx, - have hx' : x ∈ affine_span R (range b.points), + { have hx' : x ∈ affine_span R (range b), { rw b.tot, exact affine_subspace.mem_top R E x, }, obtain ⟨s, w, hw₁, rfl⟩ := (mem_affine_span_iff_eq_affine_combination R E).mp hx', refine ⟨s, w, _, hw₁, rfl⟩, diff --git a/src/analysis/inner_product_space/pi_L2.lean b/src/analysis/inner_product_space/pi_L2.lean index 866a1518b962f..6fc495ad99628 100644 --- a/src/analysis/inner_product_space/pi_L2.lean +++ b/src/analysis/inner_product_space/pi_L2.lean @@ -470,9 +470,9 @@ end ⇑(b.reindex e) = ⇑b ∘ ⇑(e.symm) := funext (b.reindex_apply e) -@[simp] protected lemma reindex_repr +@[simp] protected lemma repr_reindex (b : orthonormal_basis ι 𝕜 E) (e : ι ≃ ι') (x : E) (i' : ι') : - ((b.reindex e).repr x) i' = (b.repr x) (e.symm i') := + (b.reindex e).repr x i' = b.repr x (e.symm i') := by { classical, rw [orthonormal_basis.repr_apply_apply, b.repr_apply_apply, orthonormal_basis.coe_reindex] } diff --git a/src/analysis/normed_space/add_torsor_bases.lean b/src/analysis/normed_space/add_torsor_bases.lean index 87ed4056ec2a8..6da2223d09839 100644 --- a/src/analysis/normed_space/add_torsor_bases.lean +++ b/src/analysis/normed_space/add_torsor_bases.lean @@ -56,17 +56,17 @@ TODO Restate this result for affine spaces (instead of vector spaces) once the d convexity is generalised to this setting. -/ lemma affine_basis.interior_convex_hull {ι E : Type*} [finite ι] [normed_add_comm_group E] [normed_space ℝ E] (b : affine_basis ι ℝ E) : - interior (convex_hull ℝ (range b.points)) = {x | ∀ i, 0 < b.coord i x} := + interior (convex_hull ℝ (range b)) = {x | ∀ i, 0 < b.coord i x} := begin casesI subsingleton_or_nontrivial ι, { -- The zero-dimensional case. - have : range (b.points) = univ, + have : range b = univ, from affine_subspace.eq_univ_of_subsingleton_span_eq_top (subsingleton_range _) b.tot, simp [this] }, { -- The positive-dimensional case. haveI : finite_dimensional ℝ E := b.finite_dimensional, - have : convex_hull ℝ (range b.points) = ⋂ i, (b.coord i)⁻¹' Ici 0, - { rw [convex_hull_affine_basis_eq_nonneg_barycentric b, set_of_forall], refl }, + have : convex_hull ℝ (range b) = ⋂ i, (b.coord i)⁻¹' Ici 0, + { rw [b.convex_hull_eq_nonneg_coord, set_of_forall], refl }, ext, simp only [this, interior_Inter, ← is_open_map.preimage_interior_eq_interior_preimage (is_open_map_barycentric_coord b _) (continuous_barycentric_coord b _), @@ -132,7 +132,7 @@ top_unique $ is_open_interior.affine_span_eq_top hs ▸ (affine_span_mono _ interior_subset).trans_eq (affine_span_convex_hull _) lemma affine_basis.centroid_mem_interior_convex_hull {ι} [fintype ι] (b : affine_basis ι ℝ V) : - finset.univ.centroid ℝ b.points ∈ interior (convex_hull ℝ (range b.points)) := + finset.univ.centroid ℝ b ∈ interior (convex_hull ℝ (range b)) := begin haveI := b.nonempty, simp only [b.interior_convex_hull, mem_set_of_eq, b.coord_apply_centroid (finset.mem_univ _), @@ -144,7 +144,7 @@ lemma interior_convex_hull_nonempty_iff_affine_span_eq_top [finite_dimensional begin refine ⟨affine_span_eq_top_of_nonempty_interior, λ h, _⟩, obtain ⟨t, hts, b, hb⟩ := affine_basis.exists_affine_subbasis h, - suffices : (interior (convex_hull ℝ (range b.points))).nonempty, + suffices : (interior (convex_hull ℝ (range b))).nonempty, { rw [hb, subtype.range_coe_subtype, set_of_mem_eq] at this, refine this.mono _, mono* }, diff --git a/src/data/set/image.lean b/src/data/set/image.lean index 0b70ec820d9d6..b5d416816794a 100644 --- a/src/data/set/image.lean +++ b/src/data/set/image.lean @@ -29,13 +29,13 @@ import data.set.basic set, sets, image, preimage, pre-image, range -/ -universes u v -open function +open function set -namespace set +universes u v +variables {α β γ : Type*} {ι ι' : Sort*} -variables {α β γ : Type*} {ι : Sort*} +namespace set /-! ### Inverse image -/ @@ -914,8 +914,7 @@ end subsingleton end set namespace function - -variables {ι : Sort*} {α : Type*} {β : Type*} {f : α → β} +variables {f : α → β} open set @@ -949,7 +948,7 @@ lemma surjective.preimage_subset_preimage_iff {s t : set β} (hf : surjective f) f ⁻¹' s ⊆ f ⁻¹' t ↔ s ⊆ t := by { apply preimage_subset_preimage_iff, rw [hf.range_eq], apply subset_univ } -lemma surjective.range_comp {ι' : Sort*} {f : ι → ι'} (hf : surjective f) (g : ι' → α) : +lemma surjective.range_comp {f : ι → ι'} (hf : surjective f) (g : ι' → α) : range (g ∘ f) = range g := ext $ λ y, (@surjective.exists _ _ _ hf (λ x, g x = y)).symm @@ -981,13 +980,20 @@ by rw [← preimage_comp, h.comp_eq_id, preimage_id] end function +namespace equiv_like +variables {E : Type*} [equiv_like E ι ι'] +include ι + +@[simp] lemma range_comp (f : ι' → α) (e : E) : set.range (f ∘ e) = set.range f := +(equiv_like.surjective _).range_comp _ + +end equiv_like + /-! ### Image and preimage on subtypes -/ namespace subtype open set -variable {α : Type*} - lemma coe_image {p : α → Prop} {s : set (subtype p)} : coe '' s = {x | ∃h : p x, (⟨x, h⟩ : subtype p) ∈ s} := set.ext $ assume a, @@ -1109,9 +1115,9 @@ open function /-! ### Injectivity and surjectivity lemmas for image and preimage -/ section image_preimage -variables {α : Type u} {β : Type v} {f : α → β} -@[simp] -lemma preimage_injective : injective (preimage f) ↔ surjective f := +variables {f : α → β} + +@[simp] lemma preimage_injective : injective (preimage f) ↔ surjective f := begin refine ⟨λ h y, _, surjective.preimage_injective⟩, obtain ⟨x, hx⟩ : (f ⁻¹' {y}).nonempty, @@ -1156,7 +1162,7 @@ end set /-! ### Disjoint lemmas for image and preimage -/ section disjoint -variables {α β γ : Type*} {f : α → β} {s t : set α} +variables {f : α → β} {s t : set α} lemma disjoint.preimage (f : α → β) {s t : set β} (h : disjoint s t) : disjoint (f ⁻¹' s) (f ⁻¹' t) := diff --git a/src/linear_algebra/affine_space/basis.lean b/src/linear_algebra/affine_space/basis.lean index 56ad63b78851c..8cf3de70cf3de 100644 --- a/src/linear_algebra/affine_space/basis.lean +++ b/src/linear_algebra/affine_space/basis.lean @@ -27,7 +27,7 @@ barycentric coordinate of `q : P` is `1 - fᵢ (q -ᵥ p i)`. * `affine_basis`: a structure representing an affine basis of an affine space. * `affine_basis.coord`: the map `P →ᵃ[k] k` corresponding to `i : ι`. * `affine_basis.coord_apply_eq`: the behaviour of `affine_basis.coord i` on `p i`. - * `affine_basis.coord_apply_neq`: the behaviour of `affine_basis.coord i` on `p j` when `j ≠ i`. + * `affine_basis.coord_apply_ne`: the behaviour of `affine_basis.coord i` on `p j` when `j ≠ i`. * `affine_basis.coord_apply`: the behaviour of `affine_basis.coord i` on `p j` for general `j`. * `affine_basis.coord_apply_combination`: the characterisation of `affine_basis.coord i` in terms of affine combinations, i.e., `affine_basis.coord i (w₀ p₀ + w₁ p₁ + ⋯) = wᵢ`. @@ -44,26 +44,36 @@ open set universes u₁ u₂ u₃ u₄ /-- An affine basis is a family of affine-independent points whose span is the top subspace. -/ +@[protect_proj] structure affine_basis (ι : Type u₁) (k : Type u₂) {V : Type u₃} (P : Type u₄) [add_comm_group V] [affine_space V P] [ring k] [module k V] := -(points : ι → P) -(ind : affine_independent k points) -(tot : affine_span k (range points) = ⊤) +(to_fun : ι → P) +(ind' : affine_independent k to_fun) +(tot' : affine_span k (range to_fun) = ⊤) -variables {ι : Type u₁} {k : Type u₂} {V : Type u₃} {P : Type u₄} -variables [add_comm_group V] [affine_space V P] +variables {ι ι' k V P : Type*} [add_comm_group V] [affine_space V P] namespace affine_basis section ring -variables [ring k] [module k V] (b : affine_basis ι k P) +variables [ring k] [module k V] (b : affine_basis ι k P) {s : finset ι} {i j : ι} (e : ι ≃ ι') /-- The unique point in a single-point space is the simplest example of an affine basis. -/ instance : inhabited (affine_basis punit k punit) := -⟨{ points := id, - ind := affine_independent_of_subsingleton k id, - tot := by simp }⟩ +⟨⟨id, affine_independent_of_subsingleton k id, by simp⟩⟩ + +include V + +instance fun_like : fun_like (affine_basis ι k P) ι (λ _, P) := +{ coe := affine_basis.to_fun, + coe_injective' := λ f g h, by cases f; cases g; congr' } + +@[ext] +lemma ext {b₁ b₂ : affine_basis ι k P} (h : (b₁ : ι → P) = b₂) : b₁ = b₂ := fun_like.coe_injective h + +lemma ind : affine_independent k b := b.ind' +lemma tot : affine_span k (range b) = ⊤ := b.tot' include b @@ -72,34 +82,45 @@ not_is_empty_iff.mp $ λ hι, by simpa only [@range_eq_empty _ _ hι, affine_subspace.span_empty, bot_ne_top] using b.tot /-- Composition of an affine basis and an equivalence of index types. -/ -def comp_equiv {ι'} (e : ι' ≃ ι) : affine_basis ι' k P := -⟨b.points ∘ e, b.ind.comp_embedding e.to_embedding, by { rw [e.surjective.range_comp], exact b.3 }⟩ +def reindex (e : ι ≃ ι') : affine_basis ι' k P := +⟨b ∘ e.symm, b.ind.comp_embedding e.symm.to_embedding, + by { rw [e.symm.surjective.range_comp], exact b.3 }⟩ + +@[simp, norm_cast] lemma coe_reindex : ⇑(b.reindex e) = b ∘ e.symm := rfl +@[simp] lemma reindex_apply (i' : ι') : b.reindex e i' = b (e.symm i') := rfl + +@[simp] lemma reindex_refl : b.reindex (equiv.refl _) = b := ext rfl /-- Given an affine basis for an affine space `P`, if we single out one member of the family, we obtain a linear basis for the model space `V`. -The linear basis correpsonding to the singled-out member `i : ι` is indexed by `{j : ι // j ≠ i}` -and its `j`th element is `points j -ᵥ points i`. (See `basis_of_apply`.) -/ +The linear basis corresponding to the singled-out member `i : ι` is indexed by `{j : ι // j ≠ i}` +and its `j`th element is `b j -ᵥ b i`. (See `basis_of_apply`.) -/ noncomputable def basis_of (i : ι) : basis {j : ι // j ≠ i} k V := -basis.mk ((affine_independent_iff_linear_independent_vsub k b.points i).mp b.ind) +basis.mk ((affine_independent_iff_linear_independent_vsub k b i).mp b.ind) begin - suffices : submodule.span k (range (λ (j : {x // x ≠ i}), b.points ↑j -ᵥ b.points i)) = - vector_span k (range b.points), + suffices : submodule.span k (range (λ (j : {x // x ≠ i}), b ↑j -ᵥ b i)) = + vector_span k (range b), { rw [this, ← direction_affine_span, b.tot, affine_subspace.direction_top], exact le_rfl }, conv_rhs { rw ← image_univ, }, - rw vector_span_image_eq_span_vsub_set_right_ne k b.points (mem_univ i), + rw vector_span_image_eq_span_vsub_set_right_ne k b (mem_univ i), congr, ext v, simp, end @[simp] lemma basis_of_apply (i : ι) (j : {j : ι // j ≠ i}) : - b.basis_of i j = b.points ↑j -ᵥ b.points i := + b.basis_of i j = b ↑j -ᵥ b i := by simp [basis_of] +@[simp] lemma basis_of_reindex (i : ι') : + (b.reindex e).basis_of i = + (b.basis_of $ e.symm i).reindex (e.subtype_equiv $ λ _, e.eq_symm_apply.not) := +by { ext j, simp } + /-- The `i`th barycentric coordinate of a point. -/ noncomputable def coord (i : ι) : P →ᵃ[k] k := -{ to_fun := λ q, 1 - (b.basis_of i).sum_coords (q -ᵥ b.points i), +{ to_fun := λ q, 1 - (b.basis_of i).sum_coords (q -ᵥ b i), linear := -(b.basis_of i).sum_coords, map_vadd' := λ q v, by rw [vadd_vsub_assoc, linear_map.map_add, vadd_eq_add, linear_map.neg_apply, sub_add_eq_sub_sub_swap, add_comm, sub_eq_add_neg], } @@ -108,42 +129,41 @@ noncomputable def coord (i : ι) : P →ᵃ[k] k := (b.coord i).linear = -(b.basis_of i).sum_coords := rfl -@[simp] lemma coord_apply_eq (i : ι) : - b.coord i (b.points i) = 1 := +@[simp] lemma coord_reindex (i : ι') : + (b.reindex e).coord i = b.coord (e.symm i) := +by { ext, classical, simp [affine_basis.coord] } + +@[simp] lemma coord_apply_eq (i : ι) : b.coord i (b i) = 1 := by simp only [coord, basis.coe_sum_coords, linear_equiv.map_zero, linear_equiv.coe_coe, sub_zero, affine_map.coe_mk, finsupp.sum_zero_index, vsub_self] -@[simp] lemma coord_apply_neq (i j : ι) (h : j ≠ i) : - b.coord i (b.points j) = 0 := -by rw [coord, affine_map.coe_mk, ← subtype.coe_mk j h, ← b.basis_of_apply i ⟨j, h⟩, +@[simp] lemma coord_apply_ne (h : i ≠ j) : b.coord i (b j) = 0 := +by rw [coord, affine_map.coe_mk, ← subtype.coe_mk j h.symm, ← b.basis_of_apply, basis.sum_coords_self_apply, sub_self] -lemma coord_apply [decidable_eq ι] (i j : ι) : - b.coord i (b.points j) = if i = j then 1 else 0 := -by { cases eq_or_ne i j; simp [h.symm], simp [h], } +lemma coord_apply [decidable_eq ι] (i j : ι) : b.coord i (b j) = if i = j then 1 else 0 := +by cases eq_or_ne i j; simp [h] -@[simp] lemma coord_apply_combination_of_mem - {s : finset ι} {i : ι} (hi : i ∈ s) {w : ι → k} (hw : s.sum w = 1) : - b.coord i (s.affine_combination b.points w) = w i := +@[simp] lemma coord_apply_combination_of_mem (hi : i ∈ s) {w : ι → k} (hw : s.sum w = 1) : + b.coord i (s.affine_combination b w) = w i := begin classical, simp only [coord_apply, hi, finset.affine_combination_eq_linear_combination, if_true, mul_boole, - hw, function.comp_app, smul_eq_mul, s.sum_ite_eq, s.map_affine_combination b.points w hw], + hw, function.comp_app, smul_eq_mul, s.sum_ite_eq, s.map_affine_combination b w hw], end -@[simp] lemma coord_apply_combination_of_not_mem - {s : finset ι} {i : ι} (hi : i ∉ s) {w : ι → k} (hw : s.sum w = 1) : - b.coord i (s.affine_combination b.points w) = 0 := +@[simp] lemma coord_apply_combination_of_not_mem (hi : i ∉ s) {w : ι → k} (hw : s.sum w = 1) : + b.coord i (s.affine_combination b w) = 0 := begin classical, simp only [coord_apply, hi, finset.affine_combination_eq_linear_combination, if_false, mul_boole, - hw, function.comp_app, smul_eq_mul, s.sum_ite_eq, s.map_affine_combination b.points w hw], + hw, function.comp_app, smul_eq_mul, s.sum_ite_eq, s.map_affine_combination b w hw], end @[simp] lemma sum_coord_apply_eq_one [fintype ι] (q : P) : ∑ i, b.coord i q = 1 := begin - have hq : q ∈ affine_span k (range b.points), { rw b.tot, exact affine_subspace.mem_top k V q, }, + have hq : q ∈ affine_span k (range b), { rw b.tot, exact affine_subspace.mem_top k V q, }, obtain ⟨w, hw, rfl⟩ := eq_affine_combination_of_mem_affine_span_of_fintype hq, convert hw, ext i, @@ -151,9 +171,9 @@ begin end @[simp] lemma affine_combination_coord_eq_self [fintype ι] (q : P) : - finset.univ.affine_combination b.points (λ i, b.coord i q) = q := + finset.univ.affine_combination b (λ i, b.coord i q) = q := begin - have hq : q ∈ affine_span k (range b.points), { rw b.tot, exact affine_subspace.mem_top k V q, }, + have hq : q ∈ affine_span k (range b), { rw b.tot, exact affine_subspace.mem_top k V q, }, obtain ⟨w, hw, rfl⟩ := eq_affine_combination_of_mem_affine_span_of_fintype hq, congr, ext i, @@ -163,14 +183,15 @@ 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 := + ∑ i, b.coord i v • b 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₂ := +lemma ext_elem [finite ι] {q₁ q₂ : P} (h : ∀ i, b.coord i q₁ = b.coord i q₂) : q₁ = q₂ := begin + casesI nonempty_fintype ι, rw [← b.affine_combination_coord_eq_self q₁, ← b.affine_combination_coord_eq_self q₂], simp only [h], end @@ -179,7 +200,7 @@ end (b.coord i : P → k) = 1 := begin ext q, - have hp : (range b.points).subsingleton, + have hp : (range b).subsingleton, { rw ← image_univ, apply subsingleton.image, apply subsingleton_of_subsingleton, }, @@ -187,7 +208,7 @@ begin let s : finset ι := {i}, have hi : i ∈ s, { simp, }, have hw : s.sum (function.const ι (1 : k)) = 1, { simp, }, - have hq : q = s.affine_combination b.points (function.const ι (1 : k)), { simp, }, + have hq : q = s.affine_combination b (function.const ι (1 : k)), { simp, }, rw [pi.one_apply, hq, b.coord_apply_combination_of_mem hi hw], end @@ -202,7 +223,7 @@ begin have hj : j ∈ s, { simp, }, let w : ι → k := λ j', if j' = i then x else 1-x, have hw : s.sum w = 1, { simp [hij, finset.sum_ite, finset.filter_insert, finset.filter_eq'], }, - use s.affine_combination b.points w, + use s.affine_combination b w, simp [b.coord_apply_combination_of_mem hi hw], end @@ -231,12 +252,12 @@ include V @[simp] lemma coord_apply_centroid [char_zero k] (b : affine_basis ι k P) {s : finset ι} {i : ι} (hi : i ∈ s) : - b.coord i (s.centroid k b.points) = (s.card : k) ⁻¹ := + b.coord i (s.centroid k b) = (s.card : k) ⁻¹ := by rw [finset.centroid, b.coord_apply_combination_of_mem hi (s.sum_centroid_weights_eq_one_of_nonempty _ ⟨i, hi⟩), finset.centroid_weights] lemma exists_affine_subbasis {t : set P} (ht : affine_span k t = ⊤) : - ∃ (s ⊆ t) (b : affine_basis ↥s k P), b.points = coe := + ∃ (s ⊆ t) (b : affine_basis ↥s k P), ⇑b = coe := begin obtain ⟨s, hst, h_tot, h_ind⟩ := exists_affine_independent k V t, refine ⟨s, hst, ⟨coe, h_ind, _⟩, rfl⟩, @@ -245,7 +266,7 @@ end variables (k V P) -lemma exists_affine_basis : ∃ (s : set P) (b : affine_basis ↥s k P), b.points = coe := +lemma exists_affine_basis : ∃ (s : set P) (b : affine_basis ↥s k P), ⇑b = coe := let ⟨s, _, hs⟩ := exists_affine_subbasis (affine_subspace.span_univ k V P) in ⟨s, hs⟩ end division_ring diff --git a/src/linear_algebra/affine_space/finite_dimensional.lean b/src/linear_algebra/affine_space/finite_dimensional.lean index 568388f982953..a4934d582cc60 100644 --- a/src/linear_algebra/affine_space/finite_dimensional.lean +++ b/src/linear_algebra/affine_space/finite_dimensional.lean @@ -266,6 +266,12 @@ begin exact hi.affine_span_eq_of_le_of_card_eq_finrank_add_one le_top hc, }, end +lemma affine.simplex.span_eq_top [finite_dimensional k V] {n : ℕ} (T : affine.simplex k V n) + (hrank : finrank k V = n) : + affine_span k (set.range T.points) = ⊤ := +by rw [affine_independent.affine_span_eq_top_iff_card_eq_finrank_add_one T.independent, + fintype.card_fin, hrank] + /-- The `vector_span` of adding a point to a finite-dimensional subspace is finite-dimensional. -/ instance finite_dimensional_vector_span_insert (s : affine_subspace k P) [finite_dimensional k s.direction] (p : P) : @@ -765,7 +771,7 @@ lemma exists_affine_basis_of_finite_dimensional [fintype ι] [finite_dimensional begin obtain ⟨s, b, hb⟩ := affine_basis.exists_affine_basis k V P, lift s to finset P using b.finite_set, - refine ⟨b.comp_equiv $ fintype.equiv_of_card_eq _⟩, + refine ⟨b.reindex $ fintype.equiv_of_card_eq _⟩, rw [h, ← b.card_eq_finrank_add_one] end diff --git a/src/linear_algebra/affine_space/matrix.lean b/src/linear_algebra/affine_space/matrix.lean index b0109bc3f186c..6376fea840f17 100644 --- a/src/linear_algebra/affine_space/matrix.lean +++ b/src/linear_algebra/affine_space/matrix.lean @@ -39,7 +39,7 @@ noncomputable def to_matrix {ι' : Type*} (q : ι' → P) : matrix ι' ι k := rfl @[simp] lemma to_matrix_self [decidable_eq ι] : - b.to_matrix b.points = (1 : matrix ι ι k) := + b.to_matrix b = (1 : matrix ι ι k) := begin ext i j, rw [to_matrix_apply, coord_apply, matrix.one_eq_pi_single, pi.single_apply], @@ -74,7 +74,7 @@ coordinates of `p` with respect `b` has a left inverse, then `p` spans the entir lemma affine_span_eq_top_of_to_matrix_left_inv [decidable_eq ι] [nontrivial k] (p : ι' → P) {A : matrix ι ι' k} (hA : A ⬝ b.to_matrix p = 1) : affine_span k (range p) = ⊤ := begin - suffices : ∀ i, b.points i ∈ affine_span k (range p), + suffices : ∀ i, b i ∈ affine_span k (range p), { rw [eq_top_iff, ← b.tot, affine_span_le], rintros q ⟨i, rfl⟩, exact this i, }, @@ -85,7 +85,7 @@ begin ... = ∑ l, ∑ j, (A i j) * b.to_matrix p j l : by rw finset.sum_comm ... = ∑ l, (A ⬝ b.to_matrix p) i l : rfl ... = 1 : by simp [hA, matrix.one_apply, finset.filter_eq], }, - have hbi : b.points i = finset.univ.affine_combination p (A i), + have hbi : b i = finset.univ.affine_combination p (A i), { apply b.ext_elem, intros j, rw [b.coord_apply, finset.univ.map_affine_combination _ _ hAi, @@ -100,7 +100,7 @@ end See also `affine_basis.to_matrix_inv_mul_affine_basis_to_matrix`. -/ @[simp] lemma to_matrix_vec_mul_coords (x : P) : - (b.to_matrix b₂.points).vec_mul (b₂.coords x) = b.coords x := + (b.to_matrix b₂).vec_mul (b₂.coords x) = b.coords x := begin ext j, change _ = b.coord j x, @@ -112,17 +112,17 @@ end variables [decidable_eq ι] lemma to_matrix_mul_to_matrix : - (b.to_matrix b₂.points) ⬝ (b₂.to_matrix b.points) = 1 := + (b.to_matrix b₂) ⬝ (b₂.to_matrix b) = 1 := begin ext l m, - change (b₂.to_matrix b.points).vec_mul (b.coords (b₂.points l)) m = _, + change (b₂.to_matrix b).vec_mul (b.coords (b₂ l)) m = _, rw [to_matrix_vec_mul_coords, coords_apply, ← to_matrix_apply, to_matrix_self], end lemma is_unit_to_matrix : - is_unit (b.to_matrix b₂.points) := -⟨{ val := b.to_matrix b₂.points, - inv := b₂.to_matrix b.points, + is_unit (b.to_matrix b₂) := +⟨{ val := b.to_matrix b₂, + inv := b₂.to_matrix b, val_inv := b.to_matrix_mul_to_matrix b₂, inv_val := b₂.to_matrix_mul_to_matrix b, }, rfl⟩ @@ -136,7 +136,7 @@ begin b.affine_span_eq_top_of_to_matrix_left_inv p hA'⟩, }, { rintros ⟨h_tot, h_ind⟩, let b' : affine_basis ι k P := ⟨p, h_tot, h_ind⟩, - change is_unit (b.to_matrix b'.points), + change is_unit (b.to_matrix b'), exact b.is_unit_to_matrix b', }, end @@ -150,7 +150,7 @@ variables (b b₂ : affine_basis ι k P) See also `affine_basis.to_matrix_vec_mul_coords`. -/ @[simp] lemma to_matrix_inv_vec_mul_to_matrix (x : P) : - (b.to_matrix b₂.points)⁻¹.vec_mul (b.coords x) = b₂.coords x := + (b.to_matrix b₂)⁻¹.vec_mul (b.coords x) = b₂.coords x := begin have hu := b.is_unit_to_matrix b₂, rw matrix.is_unit_iff_is_unit_det at hu, @@ -161,7 +161,7 @@ end /-- If we fix a background affine basis `b`, then for any other basis `b₂`, we can characterise the barycentric coordinates provided by `b₂` in terms of determinants relative to `b`. -/ lemma det_smul_coords_eq_cramer_coords (x : P) : - (b.to_matrix b₂.points).det • b₂.coords x = (b.to_matrix b₂.points)ᵀ.cramer (b.coords x) := + (b.to_matrix b₂).det • b₂.coords x = (b.to_matrix b₂)ᵀ.cramer (b.coords x) := begin have hu := b.is_unit_to_matrix b₂, rw matrix.is_unit_iff_is_unit_det at hu, diff --git a/src/linear_algebra/basis.lean b/src/linear_algebra/basis.lean index 215594ec368a7..788860212e812 100644 --- a/src/linear_algebra/basis.lean +++ b/src/linear_algebra/basis.lean @@ -101,9 +101,20 @@ variables (b b₁ : basis ι R M) (i : ι) (c : R) (x : M) section repr +lemma repr_injective : injective (repr : basis ι R M → M ≃ₗ[R] (ι →₀ R)) := +λ f g h, by cases f; cases g; congr' + /-- `b i` is the `i`th basis vector. -/ -instance : has_coe_to_fun (basis ι R M) (λ _, ι → M) := -{ coe := λ b i, b.repr.symm (finsupp.single i 1) } +instance fun_like : fun_like (basis ι R M) ι (λ _, M) := +{ coe := λ b i, b.repr.symm (finsupp.single i 1), + coe_injective' := λ f g h, repr_injective $ linear_equiv.symm_bijective.injective begin + ext x, + rw [←finsupp.sum_single x, map_finsupp_sum, map_finsupp_sum], + congr' with i r, + have := congr_fun h i, + dsimp at this, + rw [←mul_one r, ←finsupp.smul_single', linear_equiv.map_smul, linear_equiv.map_smul, this], + end } @[simp] lemma coe_of_repr (e : M ≃ₗ[R] (ι →₀ R)) : ⇑(of_repr e) = λ i, e.symm (finsupp.single i 1) := @@ -275,15 +286,11 @@ begin end /-- Two bases are equal if they assign the same coordinates. -/ -lemma eq_of_repr_eq_repr {b₁ b₂ : basis ι R M} (h : ∀ x i, b₁.repr x i = b₂.repr x i) : - b₁ = b₂ := -have b₁.repr = b₂.repr, by { ext, apply h }, -by { cases b₁, cases b₂, simpa } +lemma eq_of_repr_eq_repr {b₁ b₂ : basis ι R M} (h : ∀ x i, b₁.repr x i = b₂.repr x i) : b₁ = b₂ := +repr_injective $ by { ext, apply h } /-- Two bases are equal if their basis vectors are the same. -/ -@[ext] lemma eq_of_apply_eq {b₁ b₂ : basis ι R M} (h : ∀ i, b₁ i = b₂ i) : b₁ = b₂ := -suffices b₁.repr = b₂.repr, by { cases b₁, cases b₂, simpa }, -repr_eq_iff'.mpr (λ i, by rw [h, b₂.repr_self]) +@[ext] lemma eq_of_apply_eq {b₁ b₂ : basis ι R M} : (∀ i, b₁ i = b₂ i) → b₁ = b₂ := fun_like.ext _ _ end ext @@ -347,23 +354,25 @@ by rw [linear_equiv.symm_trans_apply, finsupp.dom_lcongr_symm, finsupp.dom_lcong @[simp] lemma coe_reindex : (b.reindex e : ι' → M) = b ∘ e.symm := funext (b.reindex_apply e) -@[simp] lemma coe_reindex_repr : ((b.reindex e).repr x : ι' → R) = b.repr x ∘ e.symm := -funext $ λ i', -show (finsupp.dom_lcongr e : _ ≃ₗ[R] _) (b.repr x) i' = _, -by simp +lemma repr_reindex_apply (i' : ι') : (b.reindex e).repr x i' = b.repr x (e.symm i') := +show (finsupp.dom_lcongr e : _ ≃ₗ[R] _) (b.repr x) i' = _, by simp -@[simp] lemma reindex_repr (i' : ι') : (b.reindex e).repr x i' = b.repr x (e.symm i') := -by rw coe_reindex_repr +@[simp] lemma repr_reindex : (b.reindex e).repr x = (b.repr x).map_domain e := +fun_like.ext _ _ $ by simp [repr_reindex_apply] @[simp] lemma reindex_refl : b.reindex (equiv.refl ι) = b := eq_of_apply_eq $ λ i, by simp -/-- `simp` normal form version of `range_reindex` -/ -@[simp] lemma range_reindex' : set.range (b ∘ e.symm) = set.range b := -by rw [range_comp, equiv.range_eq_univ, set.image_univ] - +/-- `simp` can prove this as `basis.coe_reindex` + `equiv_like.range_comp` -/ lemma range_reindex : set.range (b.reindex e) = set.range b := -by rw [coe_reindex, range_reindex'] +by rw [coe_reindex, equiv_like.range_comp] + +@[simp] lemma sum_coords_reindex : (b.reindex e).sum_coords = b.sum_coords := +begin + ext x, + simp only [coe_sum_coords, repr_reindex], + exact finsupp.sum_map_domain_index (λ _, rfl) (λ _ _ _, rfl), +end /-- `b.reindex_range` is a basis indexed by `range b`, the basis vectors themselves. -/ def reindex_range : basis (range b) R M := @@ -441,7 +450,7 @@ lemma reindex_finset_range_repr_self (i : ι) : finsupp.single ⟨b i, finset.mem_image_of_mem b (finset.mem_univ i)⟩ 1 := begin ext ⟨bi, hbi⟩, - rw [reindex_finset_range, reindex_repr, reindex_range_repr_self], + rw [reindex_finset_range, repr_reindex, finsupp.map_domain_equiv_apply, reindex_range_repr_self], convert finsupp.single_apply_left ((equiv.refl M).subtype_equiv _).symm.injective _ _ _, refl end diff --git a/src/linear_algebra/matrix/basis.lean b/src/linear_algebra/matrix/basis.lean index 6064a764e2492..bd9c8e5356b06 100644 --- a/src/linear_algebra/matrix/basis.lean +++ b/src/linear_algebra/matrix/basis.lean @@ -214,8 +214,9 @@ by { haveI := classical.dec_eq ι', lemma basis.to_matrix_reindex' [decidable_eq ι] [decidable_eq ι'] (b : basis ι R M) (v : ι' → M) (e : ι ≃ ι') : (b.reindex e).to_matrix v = matrix.reindex_alg_equiv _ e (b.to_matrix (v ∘ e)) := -by { ext, simp only [basis.to_matrix_apply, basis.reindex_repr, matrix.reindex_alg_equiv_apply, - matrix.reindex_apply, matrix.submatrix_apply, function.comp_app, e.apply_symm_apply] } +by { ext, simp only [basis.to_matrix_apply, basis.repr_reindex, matrix.reindex_alg_equiv_apply, + matrix.reindex_apply, matrix.submatrix_apply, function.comp_app, e.apply_symm_apply, + finsupp.map_domain_equiv_apply] } end fintype @@ -244,7 +245,8 @@ matrix.invertible_of_left_inverse _ _ (basis.to_matrix_mul_to_matrix_flip _ _) lemma basis.to_matrix_reindex (b : basis ι R M) (v : ι' → M) (e : ι ≃ ι') : (b.reindex e).to_matrix v = (b.to_matrix v).submatrix e.symm id := -by { ext, simp only [basis.to_matrix_apply, basis.reindex_repr, matrix.submatrix_apply, id.def] } +by { ext, simp only [basis.to_matrix_apply, basis.repr_reindex, matrix.submatrix_apply, id.def, + finsupp.map_domain_equiv_apply] } @[simp] lemma basis.to_matrix_map (b : basis ι R M) (f : M ≃ₗ[R] N) (v : ι → N) : diff --git a/src/linear_algebra/matrix/spectrum.lean b/src/linear_algebra/matrix/spectrum.lean index dcb3368c8c7b4..57f56f3598dcc 100644 --- a/src/linear_algebra/matrix/spectrum.lean +++ b/src/linear_algebra/matrix/spectrum.lean @@ -99,13 +99,13 @@ begin linear_map.coe_single, pi_Lp.equiv_symm_single, linear_equiv.symm_symm, eigenvector_basis, to_lin'_apply], simp only [basis.to_matrix, basis.coe_to_orthonormal_basis_repr, basis.equiv_fun_apply], - simp_rw [orthonormal_basis.coe_to_basis_repr_apply, orthonormal_basis.reindex_repr, + simp_rw [orthonormal_basis.coe_to_basis_repr_apply, orthonormal_basis.repr_reindex, linear_equiv.symm_symm, pi_Lp.linear_equiv_apply, pi_Lp.equiv_single, mul_vec_single, mul_one], refl }, { simp only [diagonal_mul, (∘), eigenvalues, eigenvector_basis], rw [basis.to_matrix_apply, - orthonormal_basis.coe_to_basis_repr_apply, orthonormal_basis.reindex_repr, + orthonormal_basis.coe_to_basis_repr_apply, orthonormal_basis.repr_reindex, eigenvalues₀, pi_Lp.basis_fun_apply, pi_Lp.equiv_symm_single] } end From d78597269638367c3863d40d45108f52207e03cf Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 1 Feb 2023 12:38:07 +0000 Subject: [PATCH 0402/1291] refactor(data/fintype/basic): move fin_choice to a new file (#18337) There is a refactor in the works for these definitions; it will be easier to review and port the refactor if we move this all to a new file first and just forward-port the deletion. --- src/data/fintype/basic.lean | 67 -------------------------- src/data/fintype/quotient.lean | 85 +++++++++++++++++++++++++++++++++ src/model_theory/quotients.lean | 2 +- 3 files changed, 86 insertions(+), 68 deletions(-) create mode 100644 src/data/fintype/quotient.lean diff --git a/src/data/fintype/basic.lean b/src/data/fintype/basic.lean index 94075464d09fc..88ddb700bf59b 100644 --- a/src/data/fintype/basic.lean +++ b/src/data/fintype/basic.lean @@ -793,73 +793,6 @@ lemma mem_image_univ_iff_mem_range b ∈ univ.image f ↔ b ∈ set.range f := by simp -/-- An auxiliary function for `quotient.fin_choice`. Given a -collection of setoids indexed by a type `ι`, a (finite) list `l` of -indices, and a function that for each `i ∈ l` gives a term of the -corresponding quotient type, then there is a corresponding term in the -quotient of the product of the setoids indexed by `l`. -/ -def quotient.fin_choice_aux {ι : Type*} [decidable_eq ι] - {α : ι → Type*} [S : ∀ i, setoid (α i)] : - Π (l : list ι), (Π i ∈ l, quotient (S i)) → @quotient (Π i ∈ l, α i) (by apply_instance) -| [] f := ⟦λ i, false.elim⟧ -| (i :: l) f := begin - refine quotient.lift_on₂ (f i (list.mem_cons_self _ _)) - (quotient.fin_choice_aux l (λ j h, f j (list.mem_cons_of_mem _ h))) - _ _, - exact λ a l, ⟦λ j h, - if e : j = i then by rw e; exact a else - l _ (h.resolve_left e)⟧, - refine λ a₁ l₁ a₂ l₂ h₁ h₂, quotient.sound (λ j h, _), - by_cases e : j = i; simp [e], - { subst j, exact h₁ }, - { exact h₂ _ _ } -end - -theorem quotient.fin_choice_aux_eq {ι : Type*} [decidable_eq ι] - {α : ι → Type*} [S : ∀ i, setoid (α i)] : - ∀ (l : list ι) (f : Π i ∈ l, α i), quotient.fin_choice_aux l (λ i h, ⟦f i h⟧) = ⟦f⟧ -| [] f := quotient.sound (λ i h, h.elim) -| (i :: l) f := begin - simp [quotient.fin_choice_aux, quotient.fin_choice_aux_eq l], - refine quotient.sound (λ j h, _), - by_cases e : j = i; simp [e], - subst j, refl -end - -/-- Given a collection of setoids indexed by a fintype `ι` and a -function that for each `i : ι` gives a term of the corresponding -quotient type, then there is corresponding term in the quotient of the -product of the setoids. -/ -def quotient.fin_choice {ι : Type*} [decidable_eq ι] [fintype ι] - {α : ι → Type*} [S : ∀ i, setoid (α i)] - (f : Π i, quotient (S i)) : @quotient (Π i, α i) (by apply_instance) := -quotient.lift_on (@quotient.rec_on _ _ (λ l : multiset ι, - @quotient (Π i ∈ l, α i) (by apply_instance)) - finset.univ.1 - (λ l, quotient.fin_choice_aux l (λ i _, f i)) - (λ a b h, begin - have := λ a, quotient.fin_choice_aux_eq a (λ i h, quotient.out (f i)), - simp [quotient.out_eq] at this, - simp [this], - let g := λ a:multiset ι, ⟦λ (i : ι) (h : i ∈ a), quotient.out (f i)⟧, - refine eq_of_heq ((eq_rec_heq _ _).trans (_ : g a == g b)), - congr' 1, exact quotient.sound h, - end)) - (λ f, ⟦λ i, f i (finset.mem_univ _)⟧) - (λ a b h, quotient.sound $ λ i, h _ _) - -theorem quotient.fin_choice_eq {ι : Type*} [decidable_eq ι] [fintype ι] - {α : ι → Type*} [∀ i, setoid (α i)] - (f : Π i, α i) : quotient.fin_choice (λ i, ⟦f i⟧) = ⟦f⟧ := -begin - let q, swap, change quotient.lift_on q _ _ = _, - have : q = ⟦λ i h, f i⟧, - { dsimp only [q], - exact quotient.induction_on - (@finset.univ ι _).1 (λ l, quotient.fin_choice_aux_eq _ _) }, - simp [this], exact setoid.refl _ -end - namespace fintype section choose diff --git a/src/data/fintype/quotient.lean b/src/data/fintype/quotient.lean new file mode 100644 index 0000000000000..09fe4cf45be7d --- /dev/null +++ b/src/data/fintype/quotient.lean @@ -0,0 +1,85 @@ +/- +Copyright (c) 2018 Mario Carneiro. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Mario Carneiro +-/ +import data.fintype.basic + +/-! +# Quotients of families indexed by a finite type + +This file provides `quotient.fin_choice`, a mechanism to go from a finite family of quotients +to a quotient of finite families. + +## Main definitions + +* `quotient.fin_choice` + +-/ + +/-- An auxiliary function for `quotient.fin_choice`. Given a +collection of setoids indexed by a type `ι`, a (finite) list `l` of +indices, and a function that for each `i ∈ l` gives a term of the +corresponding quotient type, then there is a corresponding term in the +quotient of the product of the setoids indexed by `l`. -/ +def quotient.fin_choice_aux {ι : Type*} [decidable_eq ι] + {α : ι → Type*} [S : ∀ i, setoid (α i)] : + Π (l : list ι), (Π i ∈ l, quotient (S i)) → @quotient (Π i ∈ l, α i) (by apply_instance) +| [] f := ⟦λ i, false.elim⟧ +| (i :: l) f := begin + refine quotient.lift_on₂ (f i (list.mem_cons_self _ _)) + (quotient.fin_choice_aux l (λ j h, f j (list.mem_cons_of_mem _ h))) + _ _, + exact λ a l, ⟦λ j h, + if e : j = i then by rw e; exact a else + l _ (h.resolve_left e)⟧, + refine λ a₁ l₁ a₂ l₂ h₁ h₂, quotient.sound (λ j h, _), + by_cases e : j = i; simp [e], + { subst j, exact h₁ }, + { exact h₂ _ _ } +end + +theorem quotient.fin_choice_aux_eq {ι : Type*} [decidable_eq ι] + {α : ι → Type*} [S : ∀ i, setoid (α i)] : + ∀ (l : list ι) (f : Π i ∈ l, α i), quotient.fin_choice_aux l (λ i h, ⟦f i h⟧) = ⟦f⟧ +| [] f := quotient.sound (λ i h, h.elim) +| (i :: l) f := begin + simp [quotient.fin_choice_aux, quotient.fin_choice_aux_eq l], + refine quotient.sound (λ j h, _), + by_cases e : j = i; simp [e], + subst j, refl +end + +/-- Given a collection of setoids indexed by a fintype `ι` and a +function that for each `i : ι` gives a term of the corresponding +quotient type, then there is corresponding term in the quotient of the +product of the setoids. -/ +def quotient.fin_choice {ι : Type*} [decidable_eq ι] [fintype ι] + {α : ι → Type*} [S : ∀ i, setoid (α i)] + (f : Π i, quotient (S i)) : @quotient (Π i, α i) (by apply_instance) := +quotient.lift_on (@quotient.rec_on _ _ (λ l : multiset ι, + @quotient (Π i ∈ l, α i) (by apply_instance)) + finset.univ.1 + (λ l, quotient.fin_choice_aux l (λ i _, f i)) + (λ a b h, begin + have := λ a, quotient.fin_choice_aux_eq a (λ i h, quotient.out (f i)), + simp [quotient.out_eq] at this, + simp [this], + let g := λ a:multiset ι, ⟦λ (i : ι) (h : i ∈ a), quotient.out (f i)⟧, + refine eq_of_heq ((eq_rec_heq _ _).trans (_ : g a == g b)), + congr' 1, exact quotient.sound h, + end)) + (λ f, ⟦λ i, f i (finset.mem_univ _)⟧) + (λ a b h, quotient.sound $ λ i, h _ _) + +theorem quotient.fin_choice_eq {ι : Type*} [decidable_eq ι] [fintype ι] + {α : ι → Type*} [∀ i, setoid (α i)] + (f : Π i, α i) : quotient.fin_choice (λ i, ⟦f i⟧) = ⟦f⟧ := +begin + let q, swap, change quotient.lift_on q _ _ = _, + have : q = ⟦λ i h, f i⟧, + { dsimp only [q], + exact quotient.induction_on + (@finset.univ ι _).1 (λ l, quotient.fin_choice_aux_eq _ _) }, + simp [this], exact setoid.refl _ +end diff --git a/src/model_theory/quotients.lean b/src/model_theory/quotients.lean index f29518061d4c4..60ccfe0151a5e 100644 --- a/src/model_theory/quotients.lean +++ b/src/model_theory/quotients.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Aaron Anderson. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Aaron Anderson -/ -import data.fintype.basic +import data.fintype.quotient import model_theory.semantics /-! From 11b92770e4d49ff3982504c4dab918ac0887fe33 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 1 Feb 2023 12:38:08 +0000 Subject: [PATCH 0403/1291] feat(linear_algebra/bilinear_form): arithmetic lemmas about symm/refl/alt (#18340) The only vaguelly non-trivial result here is `bilin_form.is_refl.smul`, the rest follow essentially by definition. --- src/linear_algebra/bilinear_form.lean | 70 +++++++++++++++++++- src/linear_algebra/quadratic_form/basic.lean | 4 ++ 2 files changed, 72 insertions(+), 2 deletions(-) diff --git a/src/linear_algebra/bilinear_form.lean b/src/linear_algebra/bilinear_form.lean index eb4d411cfb3bf..1f4d13d0f31ab 100644 --- a/src/linear_algebra/bilinear_form.lean +++ b/src/linear_algebra/bilinear_form.lean @@ -673,6 +673,8 @@ end end basis +/-! ### Reflexivity, symmetry, and alternativity -/ + /-- The proposition that a bilinear form is reflexive -/ def is_refl (B : bilin_form R M) : Prop := ∀ (x y : M), B x y = 0 → B y x = 0 @@ -685,8 +687,26 @@ lemma eq_zero : ∀ {x y : M}, B x y = 0 → B y x = 0 := λ x y, H x y lemma ortho_comm {x y : M} : is_ortho B x y ↔ is_ortho B y x := ⟨eq_zero H, eq_zero H⟩ +protected lemma neg {B : bilin_form R₁ M₁} (hB : B.is_refl) : (-B).is_refl := +λ x y, neg_eq_zero.mpr ∘ hB x y ∘ neg_eq_zero.mp + +protected lemma smul {α} [semiring α] [module α R] [smul_comm_class α R R] + [no_zero_smul_divisors α R] (a : α) {B : bilin_form R M} (hB : B.is_refl) : (a • B).is_refl := +λ x y h, (smul_eq_zero.mp h).elim + (λ ha, smul_eq_zero_of_left ha _) + (λ hBz, smul_eq_zero_of_right _ (hB _ _ hBz)) + +protected lemma group_smul {α} [group α] [distrib_mul_action α R] [smul_comm_class α R R] + (a : α) {B : bilin_form R M} (hB : B.is_refl) : (a • B).is_refl := +λ x y, (smul_eq_zero_iff_eq _).mpr ∘ hB x y ∘ (smul_eq_zero_iff_eq _).mp + end is_refl +@[simp] lemma is_refl_zero : (0 : bilin_form R M).is_refl := λ _ _ _, rfl + +@[simp] lemma is_refl_neg {B : bilin_form R₁ M₁} : (-B).is_refl ↔ B.is_refl := +⟨λ h, neg_neg B ▸ h.neg, is_refl.neg⟩ + /-- The proposition that a bilinear form is symmetric -/ def is_symm (B : bilin_form R M) : Prop := ∀ (x y : M), B x y = B y x @@ -701,8 +721,30 @@ lemma is_refl : B.is_refl := λ x y H1, H x y ▸ H1 lemma ortho_comm {x y : M} : is_ortho B x y ↔ is_ortho B y x := H.is_refl.ortho_comm +protected lemma add {B₁ B₂ : bilin_form R M} (hB₁ : B₁.is_symm) (hB₂ : B₂.is_symm) : + (B₁ + B₂).is_symm := +λ x y, (congr_arg2 (+) (hB₁ x y) (hB₂ x y) : _) + +protected lemma sub {B₁ B₂ : bilin_form R₁ M₁} (hB₁ : B₁.is_symm) (hB₂ : B₂.is_symm) : + (B₁ - B₂).is_symm := +λ x y, (congr_arg2 has_sub.sub (hB₁ x y) (hB₂ x y) : _) + +protected lemma neg {B : bilin_form R₁ M₁} (hB : B.is_symm) : + (-B).is_symm := +λ x y, congr_arg has_neg.neg (hB x y) + +protected lemma smul {α} [monoid α] [distrib_mul_action α R] [smul_comm_class α R R] + (a : α) {B : bilin_form R M} (hB : B.is_symm) : + (a • B).is_symm := +λ x y, congr_arg ((•) a) (hB x y) + end is_symm +@[simp] lemma is_symm_zero : (0 : bilin_form R M).is_symm := λ _ _, rfl + +@[simp] lemma is_symm_neg {B : bilin_form R₁ M₁} : (-B).is_symm ↔ B.is_symm := +⟨λ h, neg_neg B ▸ h.neg, is_symm.neg⟩ + lemma is_symm_iff_flip' [algebra R₂ R] : B.is_symm ↔ flip_hom R₂ B = B := begin split, @@ -721,7 +763,7 @@ namespace is_alt lemma self_eq_zero (H : B.is_alt) (x : M) : B x x = 0 := H x -lemma neg (H : B₁.is_alt) (x y : M₁) : +lemma neg_eq (H : B₁.is_alt) (x y : M₁) : - B₁ x y = B₁ y x := begin have H1 : B₁ (x + y) (x + y) = 0, @@ -735,14 +777,38 @@ end lemma is_refl (H : B₁.is_alt) : B₁.is_refl := begin intros x y h, - rw [←neg H, h, neg_zero], + rw [←neg_eq H, h, neg_zero], end lemma ortho_comm (H : B₁.is_alt) {x y : M₁} : is_ortho B₁ x y ↔ is_ortho B₁ y x := H.is_refl.ortho_comm +protected lemma add {B₁ B₂ : bilin_form R M} (hB₁ : B₁.is_alt) (hB₂ : B₂.is_alt) : + (B₁ + B₂).is_alt := +λ x, (congr_arg2 (+) (hB₁ x) (hB₂ x) : _).trans $ add_zero _ + +protected lemma sub {B₁ B₂ : bilin_form R₁ M₁} (hB₁ : B₁.is_alt) (hB₂ : B₂.is_alt) : + (B₁ - B₂).is_alt := +λ x, (congr_arg2 has_sub.sub (hB₁ x) (hB₂ x)).trans $ sub_zero _ + +protected lemma neg {B : bilin_form R₁ M₁} (hB : B.is_alt) : + (-B).is_alt := +λ x, neg_eq_zero.mpr $ hB x + +protected lemma smul {α} [monoid α] [distrib_mul_action α R] [smul_comm_class α R R] + (a : α) {B : bilin_form R M} (hB : B.is_alt) : + (a • B).is_alt := +λ x, (congr_arg ((•) a) (hB x)).trans $ smul_zero _ + end is_alt +@[simp] lemma is_alt_zero : (0 : bilin_form R M).is_alt := λ _, rfl + +@[simp] lemma is_alt_neg {B : bilin_form R₁ M₁} : (-B).is_alt ↔ B.is_alt := +⟨λ h, neg_neg B ▸ h.neg, is_alt.neg⟩ + +/-! ### Linear adjoints -/ + section linear_adjoints variables (B) (F : bilin_form R M) diff --git a/src/linear_algebra/quadratic_form/basic.lean b/src/linear_algebra/quadratic_form/basic.lean index bea24f88c4f4d..61d4c6bda1277 100644 --- a/src/linear_algebra/quadratic_form/basic.lean +++ b/src/linear_algebra/quadratic_form/basic.lean @@ -576,6 +576,10 @@ map_multiset_sum (to_quadratic_form_add_monoid_hom R M) B (∑ i in s, B i).to_quadratic_form = ∑ i in s, (B i).to_quadratic_form := map_sum (to_quadratic_form_add_monoid_hom R M) B s +@[simp] lemma to_quadratic_form_eq_zero {B : bilin_form R M} : + B.to_quadratic_form = 0 ↔ B.is_alt := +quadratic_form.ext_iff + end semiring section ring From ae690b0c236e488a0043f6faa8ce3546e7f2f9c5 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 1 Feb 2023 12:38:09 +0000 Subject: [PATCH 0404/1291] refactor(linear_algebra/{general,special}_linear_group): move to matrix subfolder (#18344) This frees up `linear_algebra/general_linear_group` to be about `linear_map.general_linear_group`. --- src/analysis/complex/isometry.lean | 2 +- src/analysis/complex/upper_half_plane/basic.lean | 4 ++-- src/linear_algebra/{ => matrix}/general_linear_group.lean | 2 +- src/linear_algebra/{ => matrix}/special_linear_group.lean | 0 src/number_theory/modular.lean | 2 +- src/number_theory/modular_forms/congruence_subgroups.lean | 4 ++-- src/number_theory/modular_forms/slash_actions.lean | 4 ++-- 7 files changed, 9 insertions(+), 9 deletions(-) rename src/linear_algebra/{ => matrix}/general_linear_group.lean (99%) rename src/linear_algebra/{ => matrix}/special_linear_group.lean (100%) diff --git a/src/analysis/complex/isometry.lean b/src/analysis/complex/isometry.lean index 7718c6c229a3b..c20b7eb7f6ec9 100644 --- a/src/analysis/complex/isometry.lean +++ b/src/analysis/complex/isometry.lean @@ -5,7 +5,7 @@ Authors: François Sunatori -/ import analysis.complex.circle import linear_algebra.determinant -import linear_algebra.general_linear_group +import linear_algebra.matrix.general_linear_group /-! # Isometries of the Complex Plane diff --git a/src/analysis/complex/upper_half_plane/basic.lean b/src/analysis/complex/upper_half_plane/basic.lean index 7b234162517b4..c1f83871e5200 100644 --- a/src/analysis/complex/upper_half_plane/basic.lean +++ b/src/analysis/complex/upper_half_plane/basic.lean @@ -4,10 +4,10 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Alex Kontorovich, Heather Macbeth, Marc Masdeu -/ import data.fintype.parity -import linear_algebra.special_linear_group +import linear_algebra.matrix.special_linear_group import analysis.complex.basic import group_theory.group_action.defs -import linear_algebra.general_linear_group +import linear_algebra.matrix.general_linear_group /-! diff --git a/src/linear_algebra/general_linear_group.lean b/src/linear_algebra/matrix/general_linear_group.lean similarity index 99% rename from src/linear_algebra/general_linear_group.lean rename to src/linear_algebra/matrix/general_linear_group.lean index 37278cb102c74..9a306aad8ad3b 100644 --- a/src/linear_algebra/general_linear_group.lean +++ b/src/linear_algebra/matrix/general_linear_group.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Birkbeck -/ import linear_algebra.matrix.nonsingular_inverse -import linear_algebra.special_linear_group +import linear_algebra.matrix.special_linear_group /-! # The General Linear group $GL(n, R)$ diff --git a/src/linear_algebra/special_linear_group.lean b/src/linear_algebra/matrix/special_linear_group.lean similarity index 100% rename from src/linear_algebra/special_linear_group.lean rename to src/linear_algebra/matrix/special_linear_group.lean diff --git a/src/number_theory/modular.lean b/src/number_theory/modular.lean index bf0d3051973be..3a4f23d563c54 100644 --- a/src/number_theory/modular.lean +++ b/src/number_theory/modular.lean @@ -6,7 +6,7 @@ Authors: Alex Kontorovich, Heather Macbeth, Marc Masdeu import analysis.complex.upper_half_plane.basic import analysis.normed_space.finite_dimension -import linear_algebra.general_linear_group +import linear_algebra.matrix.general_linear_group /-! # The action of the modular group SL(2, ℤ) on the upper half-plane diff --git a/src/number_theory/modular_forms/congruence_subgroups.lean b/src/number_theory/modular_forms/congruence_subgroups.lean index 3426c626336ec..18d42bbc9f7fa 100644 --- a/src/number_theory/modular_forms/congruence_subgroups.lean +++ b/src/number_theory/modular_forms/congruence_subgroups.lean @@ -3,10 +3,10 @@ Copyright (c) 2022 Chris Birkbeck. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Birkbeck -/ -import linear_algebra.special_linear_group import data.zmod.basic -import group_theory.subgroup.pointwise import group_theory.group_action.conj_act +import group_theory.subgroup.pointwise +import linear_algebra.matrix.special_linear_group /-! # Congruence subgroups diff --git a/src/number_theory/modular_forms/slash_actions.lean b/src/number_theory/modular_forms/slash_actions.lean index e572d30bb678e..bb1ce8b989a83 100644 --- a/src/number_theory/modular_forms/slash_actions.lean +++ b/src/number_theory/modular_forms/slash_actions.lean @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Birkbeck -/ import analysis.complex.upper_half_plane.basic -import linear_algebra.general_linear_group -import linear_algebra.special_linear_group +import linear_algebra.matrix.general_linear_group +import linear_algebra.matrix.special_linear_group /-! # Slash actions From 6010cf523816335f7bae7f8584cb2edaace73940 Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Wed, 1 Feb 2023 15:31:49 +0000 Subject: [PATCH 0405/1291] feat(ring_theory/dedekind_domain): lemmas for showing Dedekind domains are PIDs (#18301) This PR includes a few lemmas that prove certain Dedekind domains are PIDs. These will be useful for showing the ideal norm map is multiplicative. Co-Authored-By: Riccardo Brasca Co-authored-by: Anne Baanen Co-authored-by: Anne Baanen Co-authored-by: Junyan Xu --- src/algebra/algebra/operations.lean | 6 + src/algebra/module/submodule/bilinear.lean | 21 ++ src/ring_theory/dedekind_domain/ideal.lean | 43 ++++ src/ring_theory/dedekind_domain/pid.lean | 252 ++++++++++++++++++++ src/ring_theory/fractional_ideal.lean | 8 + src/ring_theory/ideal/basic.lean | 7 + src/ring_theory/ideal/operations.lean | 6 + src/ring_theory/principal_ideal_domain.lean | 85 ++++++- 8 files changed, 427 insertions(+), 1 deletion(-) create mode 100644 src/ring_theory/dedekind_domain/pid.lean diff --git a/src/algebra/algebra/operations.lean b/src/algebra/algebra/operations.lean index 7945baacf4230..5ed6f9175320f 100644 --- a/src/algebra/algebra/operations.lean +++ b/src/algebra/algebra/operations.lean @@ -299,6 +299,12 @@ submodule.mem_span_mul_finite_of_mem_span_mul variables {M N P} +lemma mem_span_singleton_mul {x y : A} : x ∈ span R {y} * P ↔ ∃ z ∈ P, y * z = x := +by { simp_rw [(*), map₂_span_singleton_eq_map, exists_prop], refl } + +lemma mem_mul_span_singleton {x y : A} : x ∈ P * span R {y} ↔ ∃ z ∈ P, z * y = x := +by { simp_rw [(*), map₂_span_singleton_eq_map_flip, exists_prop], refl } + /-- Sub-R-modules of an R-algebra form a semiring. -/ instance : semiring (submodule R A) := { one_mul := submodule.one_mul, diff --git a/src/algebra/module/submodule/bilinear.lean b/src/algebra/module/submodule/bilinear.lean index 7d877ca6eec76..222d53e4c9774 100644 --- a/src/algebra/module/submodule/bilinear.lean +++ b/src/algebra/module/submodule/bilinear.lean @@ -113,6 +113,10 @@ lemma map₂_eq_span_image2 (f : M →ₗ[R] N →ₗ[R] P) (p : submodule R M) map₂ f p q = span R (set.image2 (λ m n, f m n) (p : set M) (q : set N)) := by rw [← map₂_span_span, span_eq, span_eq] +lemma map₂_flip (f : M →ₗ[R] N →ₗ[R] P) (p : submodule R M) (q : submodule R N) : + map₂ f.flip q p = map₂ f p q := +by { rw [map₂_eq_span_image2, map₂_eq_span_image2, set.image2_swap], refl } + lemma map₂_supr_left (f : M →ₗ[R] N →ₗ[R] P) (s : ι → submodule R M) (t : submodule R N) : map₂ f (⨆ i, s i) t = ⨆ i, map₂ f (s i) t := begin @@ -131,4 +135,21 @@ begin simp_rw [map₂_span_span, ← span_Union, map₂_span_span, set.image2_Union_right], end +theorem map₂_span_singleton_eq_map (f : M →ₗ[R] N →ₗ[R] P) (m : M) : + map₂ f (span R {m}) = map (f m) := +begin + funext, rw map₂_eq_span_image2, apply le_antisymm, + { rw [span_le, set.image2_subset_iff], + intros x hx y hy, + obtain ⟨a, rfl⟩ := mem_span_singleton.1 hx, + rw [f.map_smul], + exact smul_mem _ a (mem_map_of_mem hy) }, + { rintro _ ⟨n, hn, rfl⟩, + exact subset_span ⟨m, n, mem_span_singleton_self m, hn, rfl⟩ }, +end + +theorem map₂_span_singleton_eq_map_flip (f : M →ₗ[R] N →ₗ[R] P) (s : submodule R M) (n : N) : + map₂ f s (span R {n}) = map (f.flip n) s := +by rw [← map₂_span_singleton_eq_map, map₂_flip] + end submodule diff --git a/src/ring_theory/dedekind_domain/ideal.lean b/src/ring_theory/dedekind_domain/ideal.lean index 5765fa9516407..0c67434479528 100644 --- a/src/ring_theory/dedekind_domain/ideal.lean +++ b/src/ring_theory/dedekind_domain/ideal.lean @@ -1313,6 +1313,49 @@ ring_equiv.trans (ideal.quot_equiv_of_eq (inf_eq_mul_of_coprime coprime).symm) (ideal.quotient_inf_equiv_quotient_prod I J coprime) +/-- **Chinese remainder theorem** for a Dedekind domain: if the ideal `I` factors as +`∏ i in s, P i ^ e i`, then `R ⧸ I` factors as `Π (i : s), R ⧸ (P i ^ e i)`. + +This is a version of `is_dedekind_domain.quotient_equiv_pi_of_prod_eq` where we restrict +the product to a finite subset `s` of a potentially infinite indexing type `ι`. +-/ +noncomputable def is_dedekind_domain.quotient_equiv_pi_of_finset_prod_eq {ι : Type*} {s : finset ι} + (I : ideal R) (P : ι → ideal R) (e : ι → ℕ) + (prime : ∀ i ∈ s, prime (P i)) (coprime : ∀ (i j ∈ s), i ≠ j → P i ≠ P j) + (prod_eq : (∏ i in s, P i ^ e i) = I) : + R ⧸ I ≃+* Π (i : s), R ⧸ (P i ^ e i) := +is_dedekind_domain.quotient_equiv_pi_of_prod_eq I (λ (i : s), P i) (λ (i : s), e i) + (λ i, prime i i.2) + (λ i j h, coprime i i.2 j j.2 (subtype.coe_injective.ne h)) + (trans (finset.prod_coe_sort s (λ i, P i ^ e i)) prod_eq) + +/-- Corollary of the Chinese remainder theorem: given elements `x i : R / P i ^ e i`, +we can choose a representative `y : R` such that `y ≡ x i (mod P i ^ e i)`.-/ +lemma is_dedekind_domain.exists_representative_mod_finset {ι : Type*} {s : finset ι} + (P : ι → ideal R) (e : ι → ℕ) + (prime : ∀ i ∈ s, prime (P i)) (coprime : ∀ (i j ∈ s), i ≠ j → P i ≠ P j) + (x : Π (i : s), R ⧸ (P i ^ e i)) : + ∃ y, ∀ i (hi : i ∈ s), ideal.quotient.mk (P i ^ e i) y = x ⟨i, hi⟩ := +begin + let f := is_dedekind_domain.quotient_equiv_pi_of_finset_prod_eq _ P e prime coprime rfl, + obtain ⟨y, rfl⟩ := f.surjective x, + obtain ⟨z, rfl⟩ := ideal.quotient.mk_surjective y, + exact ⟨z, λ i hi, rfl⟩ +end + +/-- Corollary of the Chinese remainder theorem: given elements `x i : R`, +we can choose a representative `y : R` such that `y - x i ∈ P i ^ e i`.-/ +lemma is_dedekind_domain.exists_forall_sub_mem_ideal {ι : Type*} {s : finset ι} + (P : ι → ideal R) (e : ι → ℕ) + (prime : ∀ i ∈ s, prime (P i)) (coprime : ∀ (i j ∈ s), i ≠ j → P i ≠ P j) + (x : s → R) : + ∃ y, ∀ i (hi : i ∈ s), y - x ⟨i, hi⟩ ∈ P i ^ e i := +begin + obtain ⟨y, hy⟩ := is_dedekind_domain.exists_representative_mod_finset P e prime coprime + (λ i, ideal.quotient.mk _ (x i)), + exact ⟨y, λ i hi, ideal.quotient.eq.mp (hy i hi)⟩ +end + end dedekind_domain end chinese_remainder diff --git a/src/ring_theory/dedekind_domain/pid.lean b/src/ring_theory/dedekind_domain/pid.lean new file mode 100644 index 0000000000000..fd41b7a7737fb --- /dev/null +++ b/src/ring_theory/dedekind_domain/pid.lean @@ -0,0 +1,252 @@ +/- +Copyright (c) 2023 Anne Baanen. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Anne Baanen +-/ + +import ring_theory.dedekind_domain.dvr +import ring_theory.dedekind_domain.ideal + +/-! +# Proving a Dedekind domain is a PID + +This file contains some results that we can use to show all ideals in a Dedekind domain are +principal. + +## Main results + + * `ideal.is_principal.of_finite_maximals_of_is_unit`: an invertible ideal in a commutative ring + with finitely many maximal ideals, is a principal ideal. + * `is_principal_ideal_ring.of_finite_primes`: if a Dedekind domain has finitely many prime ideals, + it is a principal ideal domain. +-/ + +variables {R : Type*} [comm_ring R] + +open ideal +open unique_factorization_monoid +open_locale big_operators +open_locale non_zero_divisors + +open unique_factorization_monoid + +/-- Let `P` be a prime ideal, `x ∈ P \ P²` and `x ∉ Q` for all prime ideals `Q ≠ P`. +Then `P` is generated by `x`. -/ +lemma ideal.eq_span_singleton_of_mem_of_not_mem_sq_of_not_mem_prime_ne + {P : ideal R} (hP : P.is_prime) [is_domain R] [is_dedekind_domain R] + {x : R} (x_mem : x ∈ P) (hxP2 : x ∉ P^2) + (hxQ : ∀ (Q : ideal R), is_prime Q → Q ≠ P → x ∉ Q) : + P = ideal.span {x} := +begin + letI := classical.dec_eq (ideal R), + have hx0 : x ≠ 0, + { rintro rfl, + exact hxP2 (zero_mem _) }, + by_cases hP0 : P = ⊥, + { unfreezingI { subst hP0 }, + simpa using hxP2 }, + have hspan0 : span ({x} : set R) ≠ ⊥ := mt ideal.span_singleton_eq_bot.mp hx0, + have span_le := (ideal.span_singleton_le_iff_mem _).mpr x_mem, + refine associated_iff_eq.mp + ((associated_iff_normalized_factors_eq_normalized_factors hP0 hspan0).mpr + (le_antisymm ((dvd_iff_normalized_factors_le_normalized_factors hP0 hspan0).mp _) _)), + { rwa [ideal.dvd_iff_le, ideal.span_singleton_le_iff_mem] }, + simp only [normalized_factors_irreducible ((ideal.prime_of_is_prime hP0 hP).irreducible), + normalize_eq, multiset.le_iff_count, multiset.count_singleton], + intros Q, + split_ifs with hQ, + { unfreezingI { subst hQ }, + refine (ideal.count_normalized_factors_eq _ _).le; + simp only [ideal.span_singleton_le_iff_mem, pow_one]; + assumption }, + by_cases hQp : is_prime Q, + { resetI, + refine (ideal.count_normalized_factors_eq _ _).le; + simp only [ideal.span_singleton_le_iff_mem, pow_one, pow_zero, one_eq_top, submodule.mem_top], + exact hxQ _ hQp hQ }, + { exact (multiset.count_eq_zero.mpr (λ hQi, hQp (is_prime_of_prime (irreducible_iff_prime.mp + (irreducible_of_normalized_factor _ hQi))))).le } +end + +lemma fractional_ideal.is_principal_of_unit_of_comap_mul_span_singleton_eq_top + {R A : Type*} [comm_ring R] [comm_ring A] [algebra R A] {S : submonoid R} [is_localization S A] + (I : (fractional_ideal S A)ˣ) {v : A} (hv : v ∈ (↑I⁻¹ : fractional_ideal S A)) + (h : submodule.comap (algebra.linear_map R A) (I * submodule.span R {v}) = ⊤) : + submodule.is_principal (I : submodule R A) := +begin + have hinv := I.mul_inv, + set J := submodule.comap (algebra.linear_map R A) (I * submodule.span R {v}), + have hJ : is_localization.coe_submodule A J = I * submodule.span R {v}, + { rw [subtype.ext_iff, fractional_ideal.coe_mul, fractional_ideal.coe_one] at hinv, + apply submodule.map_comap_eq_self, + rw [← submodule.one_eq_range, ← hinv], + exact submodule.mul_le_mul_right ((submodule.span_singleton_le_iff_mem _ _).2 hv) }, + have : (1 : A) ∈ ↑I * submodule.span R {v}, + { rw [← hJ, h, is_localization.coe_submodule_top, submodule.mem_one], + exact ⟨1, (algebra_map R _).map_one⟩ }, + obtain ⟨w, hw, hvw⟩ := submodule.mem_mul_span_singleton.1 this, + refine ⟨⟨w, _⟩⟩, + rw [← fractional_ideal.coe_span_singleton S, ← inv_inv I, eq_comm, coe_coe], + refine congr_arg coe (units.eq_inv_of_mul_eq_one_left (le_antisymm _ _)), + { apply_instance }, + { conv_rhs { rw [← hinv, mul_comm] }, + apply fractional_ideal.mul_le_mul_left (fractional_ideal.span_singleton_le_iff_mem.mpr hw) }, + { rw [fractional_ideal.one_le, ← hvw, mul_comm], + exact fractional_ideal.mul_mem_mul hv (fractional_ideal.mem_span_singleton_self _ _) } +end + +/-- +An invertible fractional ideal of a commutative ring with finitely many maximal ideals is principal. + +https://math.stackexchange.com/a/95857 -/ +theorem fractional_ideal.is_principal.of_finite_maximals_of_inv + {A : Type*} [comm_ring A] [algebra R A] {S : submonoid R} [is_localization S A] + (hS : S ≤ R⁰) (hf : {I : ideal R | I.is_maximal}.finite) + (I I' : fractional_ideal S A) (hinv : I * I' = 1) : + submodule.is_principal (I : submodule R A) := +begin + have hinv' := hinv, + rw [subtype.ext_iff, fractional_ideal.coe_mul] at hinv, + let s := hf.to_finset, + haveI := classical.dec_eq (ideal R), + have coprime : ∀ (M ∈ s) (M' ∈ s.erase M), M ⊔ M' = ⊤, + { simp_rw [finset.mem_erase, hf.mem_to_finset], + rintro M hM M' ⟨hne, hM'⟩, + exact ideal.is_maximal.coprime_of_ne hM hM' hne.symm }, + have nle : ∀ M ∈ s, ¬ (⨅ M' ∈ s.erase M, M') ≤ M := λ M hM, left_lt_sup.1 + ((hf.mem_to_finset.1 hM).ne_top.lt_top.trans_eq (ideal.sup_infi_eq_top $ coprime M hM).symm), + have : ∀ M ∈ s, ∃ (a ∈ I) (b ∈ I'), a * b ∉ is_localization.coe_submodule A M, + { intros M hM, by_contra' h, + obtain ⟨x, hx, hxM⟩ := set_like.exists_of_lt ((is_localization.coe_submodule_strict_mono + hS (hf.mem_to_finset.1 hM).ne_top.lt_top).trans_eq hinv.symm), + refine hxM (submodule.map₂_le.2 _ hx), exact h }, + choose! a ha b hb hm using this, + choose! u hu hum using λ M hM, set_like.not_le_iff_exists.1 (nle M hM), + let v := ∑ M in s, u M • b M, + have hv : v ∈ I' := submodule.sum_mem _ (λ M hM, submodule.smul_mem _ _ $ hb M hM), + refine fractional_ideal.is_principal_of_unit_of_comap_mul_span_singleton_eq_top + (units.mk_of_mul_eq_one I I' hinv') hv (of_not_not $ λ h, _), + obtain ⟨M, hM, hJM⟩ := ideal.exists_le_maximal _ h, + replace hM := hf.mem_to_finset.2 hM, + have : ∀ (a ∈ I) (b ∈ I'), ∃ c, algebra_map R _ c = a * b, + { intros a ha b hb, have hi := hinv.le, + obtain ⟨c, -, hc⟩ := hi (submodule.mul_mem_mul ha hb), + exact ⟨c, hc⟩ }, + have hmem: a M * v ∈ is_localization.coe_submodule A M, + { obtain ⟨c, hc⟩ := this _ (ha M hM) v hv, + refine is_localization.coe_submodule_mono _ hJM ⟨c, _, hc⟩, + have := submodule.mul_mem_mul (ha M hM) (submodule.mem_span_singleton_self v), + rwa ← hc at this }, + simp_rw [finset.mul_sum, mul_smul_comm] at hmem, + rw [← s.add_sum_erase _ hM, submodule.add_mem_iff_left] at hmem, + { refine hm M hM _, + obtain ⟨c, (hc : algebra_map R A c = a M * b M)⟩ := this _ (ha M hM) _ (hb M hM), + rw ← hc at hmem ⊢, + rw [algebra.smul_def, ← _root_.map_mul] at hmem, + obtain ⟨d, hdM, he⟩ := hmem, + rw is_localization.injective _ hS he at hdM, + exact submodule.mem_map_of_mem + (((hf.mem_to_finset.1 hM).is_prime.mem_or_mem hdM).resolve_left $ hum M hM) }, + { refine submodule.sum_mem _ (λ M' hM', _), + rw finset.mem_erase at hM', + obtain ⟨c, hc⟩ := this _ (ha M hM) _ (hb M' hM'.2), + rw [← hc, algebra.smul_def, ← _root_.map_mul], + specialize hu M' hM'.2, + simp_rw [ideal.mem_infi, finset.mem_erase] at hu, + exact submodule.mem_map_of_mem (M.mul_mem_right _ $ hu M ⟨hM'.1.symm, hM⟩) }, +end + +/-- An invertible ideal in a commutative ring with finitely many maximal ideals is principal. + +https://math.stackexchange.com/a/95857 -/ +theorem ideal.is_principal.of_finite_maximals_of_is_unit + (hf : {I : ideal R | I.is_maximal}.finite) + {I : ideal R} (hI : is_unit (I : fractional_ideal R⁰ (fraction_ring R))) : + I.is_principal := +(is_localization.coe_submodule_is_principal _ le_rfl).mp + (fractional_ideal.is_principal.of_finite_maximals_of_inv le_rfl hf I + (↑(hI.unit⁻¹) : fractional_ideal R⁰ (fraction_ring R)) + hI.unit.mul_inv) + +/-- A Dedekind domain is a PID if its set of primes is finite. -/ +theorem is_principal_ideal_ring.of_finite_primes [is_domain R] [is_dedekind_domain R] + (h : {I : ideal R | I.is_prime}.finite) : + is_principal_ideal_ring R := +⟨λ I, begin + obtain rfl | hI := eq_or_ne I ⊥, + { exact bot_is_principal }, + apply ideal.is_principal.of_finite_maximals_of_is_unit, + { apply h.subset, exact @ideal.is_maximal.is_prime _ _ }, + { exact is_unit_of_mul_eq_one _ _ (fractional_ideal.coe_ideal_mul_inv I hI) }, +end⟩ + +variables [is_domain R] [is_dedekind_domain R] +variables (S : Type*) [comm_ring S] [is_domain S] +variables [algebra R S] [module.free R S] [module.finite R S] +variables (p : ideal R) (hp0 : p ≠ ⊥) [is_prime p] +variables {Sₚ : Type*} [comm_ring Sₚ] [algebra S Sₚ] +variables [is_localization (algebra.algebra_map_submonoid S p.prime_compl) Sₚ] +variables [algebra R Sₚ] [is_scalar_tower R S Sₚ] +/- The first hypothesis below follows from properties of the localization but is needed for the +second, so we leave it to the user to provide (automatically). -/ +variables [is_domain Sₚ] [is_dedekind_domain Sₚ] + +include S hp0 + +/-- If `p` is a prime in the Dedekind domain `R`, `S` an extension of `R` and `Sₚ` the localization +of `S` at `p`, then all primes in `Sₚ` are factors of the image of `p` in `Sₚ`. -/ +lemma is_localization.over_prime.mem_normalized_factors_of_is_prime [decidable_eq (ideal Sₚ)] + {P : ideal Sₚ} (hP : is_prime P) (hP0 : P ≠ ⊥) : + P ∈ normalized_factors (ideal.map (algebra_map R Sₚ) p) := +begin + have non_zero_div : algebra.algebra_map_submonoid S p.prime_compl ≤ S⁰ := + map_le_non_zero_divisors_of_injective _ (no_zero_smul_divisors.algebra_map_injective _ _) + p.prime_compl_le_non_zero_divisors, + letI : algebra (localization.at_prime p) Sₚ := localization_algebra p.prime_compl S, + haveI : is_scalar_tower R (localization.at_prime p) Sₚ := is_scalar_tower.of_algebra_map_eq + (λ x, by erw [is_localization.map_eq, is_scalar_tower.algebra_map_apply R S]), + obtain ⟨pid, p', ⟨hp'0, hp'p⟩, hpu⟩ := + (discrete_valuation_ring.iff_pid_with_one_nonzero_prime (localization.at_prime p)).mp + (is_localization.at_prime.discrete_valuation_ring_of_dedekind_domain R hp0 _), + have : local_ring.maximal_ideal (localization.at_prime p) ≠ ⊥, + { rw submodule.ne_bot_iff at ⊢ hp0, + obtain ⟨x, x_mem, x_ne⟩ := hp0, + exact ⟨algebra_map _ _ x, + (is_localization.at_prime.to_map_mem_maximal_iff _ _ _).mpr x_mem, + is_localization.to_map_ne_zero_of_mem_non_zero_divisors _ p.prime_compl_le_non_zero_divisors + (mem_non_zero_divisors_of_ne_zero x_ne)⟩ }, + rw [← multiset.singleton_le, ← normalize_eq P, + ← normalized_factors_irreducible (ideal.prime_of_is_prime hP0 hP).irreducible, + ← dvd_iff_normalized_factors_le_normalized_factors hP0, dvd_iff_le, + is_scalar_tower.algebra_map_eq R (localization.at_prime p) Sₚ, ← ideal.map_map, + localization.at_prime.map_eq_maximal_ideal, ideal.map_le_iff_le_comap, + hpu (local_ring.maximal_ideal _) ⟨this, _⟩, hpu (comap _ _) ⟨_, _⟩], + { exact le_rfl }, + { have hRS : algebra.is_integral R S := is_integral_of_noetherian + (is_noetherian_of_fg_of_noetherian' module.finite.out), + exact mt (ideal.eq_bot_of_comap_eq_bot (is_integral_localization hRS)) hP0 }, + { exact ideal.comap_is_prime (algebra_map (localization.at_prime p) Sₚ) P }, + { exact (local_ring.maximal_ideal.is_maximal _).is_prime }, + { rw [ne.def, zero_eq_bot, ideal.map_eq_bot_iff_of_injective], + { assumption }, + rw is_scalar_tower.algebra_map_eq R S Sₚ, + exact (is_localization.injective Sₚ non_zero_div).comp + (no_zero_smul_divisors.algebra_map_injective _ _) }, +end + +/-- Let `p` be a prime in the Dedekind domain `R` and `S` be an integral extension of `R`, +then the localization `Sₚ` of `S` at `p` is a PID. -/ +theorem is_dedekind_domain.is_principal_ideal_ring_localization_over_prime : + is_principal_ideal_ring Sₚ := +begin + letI := classical.dec_eq (ideal Sₚ), + letI := classical.dec_pred (λ (P : ideal Sₚ), P.is_prime), + refine is_principal_ideal_ring.of_finite_primes + (set.finite.of_finset (finset.filter (λ P, P.is_prime) + ({⊥} ∪ (normalized_factors (ideal.map (algebra_map R Sₚ) p)).to_finset)) + (λ P, _)), + rw [finset.mem_filter, finset.mem_union, finset.mem_singleton, set.mem_set_of, + multiset.mem_to_finset], + exact and_iff_right_of_imp (λ hP, or_iff_not_imp_left.mpr + (is_localization.over_prime.mem_normalized_factors_of_is_prime S p hp0 hP)) +end diff --git a/src/ring_theory/fractional_ideal.lean b/src/ring_theory/fractional_ideal.lean index f24451c9a7637..1f2f0a7445444 100644 --- a/src/ring_theory/fractional_ideal.lean +++ b/src/ring_theory/fractional_ideal.lean @@ -546,6 +546,10 @@ begin apply coe_ideal_le_one }, end +@[simp] lemma one_le {I : fractional_ideal S P} : + 1 ≤ I ↔ (1 : P) ∈ I := +by rw [← coe_le_coe, coe_one, submodule.one_le, mem_coe] + variables (S P) /-- `coe_ideal_hom (S : submonoid R) P` is `coe : ideal R → fractional_ideal S P` as a ring hom -/ @@ -1109,6 +1113,10 @@ lemma mem_span_singleton_self (x : P) : variables {S} +@[simp] lemma span_singleton_le_iff_mem {x : P} {I : fractional_ideal S P} : + span_singleton S x ≤ I ↔ x ∈ I := +by rw [← coe_le_coe, coe_span_singleton, submodule.span_singleton_le_iff_mem x ↑I, mem_coe] + lemma span_singleton_eq_span_singleton [no_zero_smul_divisors R P] {x y : P} : span_singleton S x = span_singleton S y ↔ ∃ z : Rˣ, z • x = y := by { rw [← submodule.span_singleton_eq_span_singleton, span_singleton, span_singleton], diff --git a/src/ring_theory/ideal/basic.lean b/src/ring_theory/ideal/basic.lean index 9a123aa170a2f..c84f67ef4343f 100644 --- a/src/ring_theory/ideal/basic.lean +++ b/src/ring_theory/ideal/basic.lean @@ -237,6 +237,13 @@ begin exact submodule.singleton_span_is_compact_element 1, end +lemma is_maximal.coprime_of_ne {M M' : ideal α} (hM : M.is_maximal) (hM' : M'.is_maximal) + (hne : M ≠ M') : M ⊔ M' = ⊤ := +begin + contrapose! hne with h, + exact hM.eq_of_le hM'.ne_top (le_sup_left.trans_eq (hM'.eq_of_le h le_sup_right).symm) +end + /-- **Krull's theorem**: if `I` is an ideal that is not the whole ring, then it is included in some maximal ideal. -/ theorem exists_le_maximal (I : ideal α) (hI : I ≠ ⊤) : diff --git a/src/ring_theory/ideal/operations.lean b/src/ring_theory/ideal/operations.lean index cbb6eebef923e..f61bf2be7d13b 100644 --- a/src/ring_theory/ideal/operations.lean +++ b/src/ring_theory/ideal/operations.lean @@ -497,6 +497,12 @@ lemma prod_span_singleton {ι : Type*} (s : finset ι) (I : ι → R) : (∏ i in s, ideal.span ({I i} : set R)) = ideal.span {∏ i in s, I i} := submodule.prod_span_singleton s I +@[simp] lemma multiset_prod_span_singleton (m : multiset R) : + (m.map (λ x, ideal.span {x})).prod = ideal.span ({multiset.prod m} : set R) := +multiset.induction_on m (by simp) + (λ a m ih, by simp only [multiset.map_cons, multiset.prod_cons, ih, + ← ideal.span_singleton_mul_span_singleton]) + lemma finset_inf_span_singleton {ι : Type*} (s : finset ι) (I : ι → R) (hI : set.pairwise ↑s (is_coprime on I)) : (s.inf $ λ i, ideal.span ({I i} : set R)) = ideal.span {∏ i in s, I i} := diff --git a/src/ring_theory/principal_ideal_domain.lean b/src/ring_theory/principal_ideal_domain.lean index ccf8a7be9eacb..fb6fdcf722a6c 100644 --- a/src/ring_theory/principal_ideal_domain.lean +++ b/src/ring_theory/principal_ideal_domain.lean @@ -52,7 +52,7 @@ instance top_is_principal : (⊤ : submodule R R).is_principal := variables (R) /-- A ring is a principal ideal ring if all (left) ideals are principal. -/ -class is_principal_ideal_ring (R : Type u) [ring R] : Prop := +@[mk_iff] class is_principal_ideal_ring (R : Type u) [ring R] : Prop := (principal : ∀ (S : ideal R), S.is_principal) attribute [instance] is_principal_ideal_ring.principal @@ -407,3 +407,86 @@ theorem exists_associated_pow_of_mul_eq_pow' {a b c : R} exists_associated_pow_of_mul_eq_pow ((gcd_is_unit_iff _ _).mpr hab) h end + +section principal_of_prime + +open set ideal + +variables (R) [comm_ring R] + +/-- `non_principals R` is the set of all ideals of `R` that are not principal ideals. -/ +def non_principals := {I : ideal R | ¬ I.is_principal} + +lemma non_principals_def {I : ideal R} : I ∈ non_principals R ↔ ¬ I.is_principal := +iff.rfl + +variables {R} +lemma non_principals_eq_empty_iff : non_principals R = ∅ ↔ is_principal_ideal_ring R := +by simp [set.eq_empty_iff_forall_not_mem, is_principal_ideal_ring_iff, non_principals_def] + +/-- Any chain in the set of non-principal ideals has an upper bound which is non-principal. +(Namely, the union of the chain is such an upper bound.) +-/ +lemma non_principals_zorn (c : set (ideal R)) (hs : c ⊆ non_principals R) (hchain : is_chain (≤) c) + {K : ideal R} (hKmem : K ∈ c) : + ∃ I ∈ non_principals R, ∀ J ∈ c, J ≤ I := +begin + refine ⟨Sup c, _, λ J hJ, le_Sup hJ⟩, + rintro ⟨x, hx⟩, + have hxmem : x ∈ Sup c := (hx.symm ▸ submodule.mem_span_singleton_self x), + obtain ⟨J, hJc, hxJ⟩ := (submodule.mem_Sup_of_directed ⟨K, hKmem⟩ hchain.directed_on).1 hxmem, + have hSupJ : Sup c = J := le_antisymm (by simp [hx, ideal.span_le, hxJ]) (le_Sup hJc), + specialize hs hJc, + rw [← hSupJ, hx, non_principals_def] at hs, + exact hs ⟨⟨x, rfl⟩⟩ +end + +/-- If all prime ideals in a commutative ring are principal, so are all other ideals. -/ +theorem is_principal_ideal_ring.of_prime (H : ∀ (P : ideal R), P.is_prime → P.is_principal) : + is_principal_ideal_ring R := +begin + -- Suppose the set of `non_principals` is not empty. + rw [← non_principals_eq_empty_iff, set.eq_empty_iff_forall_not_mem], + intros J hJ, + -- We will show a maximal element `I ∈ non_principals R` (which exists by Zorn) is prime. + obtain ⟨I, Ibad, -, Imax⟩ := zorn_nonempty_partial_order₀ + (non_principals R) non_principals_zorn _ hJ, + have Imax' : ∀ {J}, I < J → J.is_principal, + { intros J hJ, + by_contra He, + exact hJ.ne (Imax _ ((non_principals_def R).2 He) hJ.le).symm }, + by_cases hI1 : I = ⊤, + { subst hI1, + exact Ibad top_is_principal }, + -- Let `x y : R` with `x * y ∈ I` and suppose WLOG `y ∉ I`. + refine Ibad (H I ⟨hI1, λ x y hxy, or_iff_not_imp_right.mpr (λ hy, _)⟩), + obtain ⟨a, ha⟩ : (I ⊔ span {y}).is_principal := + Imax' (left_lt_sup.mpr (mt I.span_singleton_le_iff_mem.mp hy)), + -- Then `x ∈ I.colon (span {y})`, which is equal to `I` if it's not principal. + suffices He : ¬ ((I.colon (span {y})).is_principal), + { rw ← Imax _ ((non_principals_def R).2 He) + (λ a ha, ideal.mem_colon_singleton.2 (mul_mem_right _ _ ha)), + exact ideal.mem_colon_singleton.2 hxy }, + -- So suppose for the sake of contradiction that both `I ⊔ span {y}` and `I.colon (span {y})` + -- are principal. + rintros ⟨b, hb⟩, + -- We will show `I` is generated by `a * b`. + refine (non_principals_def _).1 Ibad ⟨⟨a * b, le_antisymm (λ i hi, _) $ + (span_singleton_mul_span_singleton a b).ge.trans _⟩⟩, + { have hisup : i ∈ I ⊔ span {y} := ideal.mem_sup_left hi, + have : y ∈ I ⊔ span {y} := ideal.mem_sup_right (ideal.mem_span_singleton_self y), + erw [ha, mem_span_singleton'] at hisup this, + obtain ⟨v, rfl⟩ := this, + obtain ⟨u, rfl⟩ := hisup, + have hucolon : u ∈ I.colon (span {v * a}), + { rw [ideal.mem_colon_singleton, mul_comm v, ← mul_assoc], + exact mul_mem_right _ _ hi }, + erw [hb, mem_span_singleton'] at hucolon, + obtain ⟨z, rfl⟩ := hucolon, + exact mem_span_singleton'.2 ⟨z, by ring⟩ }, + { rw [← ideal.submodule_span_eq, ← ha, ideal.sup_mul, sup_le_iff, + span_singleton_mul_span_singleton, mul_comm y, ideal.span_singleton_le_iff_mem], + exact ⟨mul_le_right, ideal.mem_colon_singleton.1 $ hb.symm ▸ ideal.mem_span_singleton_self b⟩ }, +end + +end principal_of_prime From 938fead7abdc0cbbca8eba7a1052865a169dc102 Mon Sep 17 00:00:00 2001 From: Oliver Nash Date: Wed, 1 Feb 2023 19:20:34 +0000 Subject: [PATCH 0406/1291] fix(algebra/lie/*): use correct terminology for `centralizer` and `normalizer` in Lie theory (#18348) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The naming here was wrong. If `N : lie_submodule R L M` then: * `{ m : M | ∀ (x : L), ⁅x, m⁆ ∈ N }` is the carrier of the _normalizer_ of `N`. If `N : set M` then: * `{ x : L | ∀ (m ∈ N), ⁅x, (m : M)⁆ = 0 }` is the carrier of the _centralizer_ of `N`. [No formal definition yet, seems important not to restrict `N` to be a `lie_submodule` since an important case is to apply this to `M = L` regarded as a Lie module over itself and `N` a `lie_subalgebra` of `L` (but not necessarily a `lie_ideal`). Note this generalises `lie_module.ker`.] I'm a bit shocked I got this muddled originally. --- src/algebra/lie/cartan_subalgebra.lean | 14 ++-- src/algebra/lie/engel.lean | 2 +- src/algebra/lie/nilpotent.lean | 20 +++--- .../lie/{centralizer.lean => normalizer.lean} | 66 +++++++++---------- src/algebra/lie/weights.lean | 2 +- 5 files changed, 52 insertions(+), 52 deletions(-) rename src/algebra/lie/{centralizer.lean => normalizer.lean} (75%) diff --git a/src/algebra/lie/cartan_subalgebra.lean b/src/algebra/lie/cartan_subalgebra.lean index 296ff9f84ed19..6ef9ca921b8ff 100644 --- a/src/algebra/lie/cartan_subalgebra.lean +++ b/src/algebra/lie/cartan_subalgebra.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Oliver Nash -/ import algebra.lie.nilpotent -import algebra.lie.centralizer +import algebra.lie.normalizer /-! # Cartan subalgebras @@ -47,10 +47,10 @@ class is_cartan_subalgebra : Prop := instance [H.is_cartan_subalgebra] : lie_algebra.is_nilpotent R H := is_cartan_subalgebra.nilpotent -@[simp] lemma centralizer_eq_self_of_is_cartan_subalgebra +@[simp] lemma normalizer_eq_self_of_is_cartan_subalgebra (H : lie_subalgebra R L) [H.is_cartan_subalgebra] : - H.to_lie_submodule.centralizer = H.to_lie_submodule := -by rw [← lie_submodule.coe_to_submodule_eq_iff, coe_centralizer_eq_normalizer, + H.to_lie_submodule.normalizer = H.to_lie_submodule := +by rw [← lie_submodule.coe_to_submodule_eq_iff, coe_normalizer_eq_normalizer, is_cartan_subalgebra.self_normalizing, coe_to_lie_submodule] @[simp] lemma ucs_eq_self_of_is_cartan_subalgebra @@ -70,8 +70,8 @@ begin have h₁ : _root_.lie_algebra.is_nilpotent R H := by apply_instance, obtain ⟨k, hk⟩ := H.to_lie_submodule.is_nilpotent_iff_exists_self_le_ucs.mp h₁, replace hk : H.to_lie_submodule = lie_submodule.ucs k ⊥ := - le_antisymm hk (lie_submodule.ucs_le_of_centralizer_eq_self - H.centralizer_eq_self_of_is_cartan_subalgebra k), + le_antisymm hk (lie_submodule.ucs_le_of_normalizer_eq_self + H.normalizer_eq_self_of_is_cartan_subalgebra k), refine ⟨k, λ l hl, _⟩, rw [← nat.sub_add_cancel hl, lie_submodule.ucs_add, ← hk, lie_subalgebra.ucs_eq_self_of_is_cartan_subalgebra], }, @@ -90,7 +90,7 @@ begin have hk' := hk (k + 1) k.le_succ, rw [lie_submodule.ucs_succ, hk k (le_refl k)] at hk', rw [← lie_subalgebra.coe_to_submodule_eq_iff, - ← lie_subalgebra.coe_centralizer_eq_normalizer, hk', + ← lie_subalgebra.coe_normalizer_eq_normalizer, hk', lie_subalgebra.coe_to_lie_submodule], end } }, end diff --git a/src/algebra/lie/engel.lean b/src/algebra/lie/engel.lean index 78ba41c33c521..fa9f351cae10e 100644 --- a/src/algebra/lie/engel.lean +++ b/src/algebra/lie/engel.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Oliver Nash -/ import algebra.lie.nilpotent -import algebra.lie.centralizer +import algebra.lie.normalizer /-! # Engel's theorem diff --git a/src/algebra/lie/nilpotent.lean b/src/algebra/lie/nilpotent.lean index acf0c6fba770b..3f7c43ddd61d1 100644 --- a/src/algebra/lie/nilpotent.lean +++ b/src/algebra/lie/nilpotent.lean @@ -5,7 +5,7 @@ Authors: Oliver Nash -/ import algebra.lie.solvable import algebra.lie.quotient -import algebra.lie.centralizer +import algebra.lie.normalizer import linear_algebra.eigenspace import ring_theory.nilpotent @@ -350,17 +350,17 @@ variables {N₁ N₂ : lie_submodule R L M} See also `lie_submodule.lcs`. -/ def ucs (k : ℕ) : lie_submodule R L M → lie_submodule R L M := -centralizer^[k] +normalizer^[k] @[simp] lemma ucs_zero : N.ucs 0 = N := rfl @[simp] lemma ucs_succ (k : ℕ) : - N.ucs (k + 1) = (N.ucs k).centralizer := -function.iterate_succ_apply' centralizer k N + N.ucs (k + 1) = (N.ucs k).normalizer := +function.iterate_succ_apply' normalizer k N lemma ucs_add (k l : ℕ) : N.ucs (k + l) = (N.ucs l).ucs k := -function.iterate_add_apply centralizer k l N +function.iterate_add_apply normalizer k l N @[mono] lemma ucs_mono (k : ℕ) (h : N₁ ≤ N₂) : N₁.ucs k ≤ N₂.ucs k := @@ -370,18 +370,18 @@ begin mono, end -lemma ucs_eq_self_of_centralizer_eq_self (h : N₁.centralizer = N₁) (k : ℕ) : +lemma ucs_eq_self_of_normalizer_eq_self (h : N₁.normalizer = N₁) (k : ℕ) : N₁.ucs k = N₁ := by { induction k with k ih, { simp, }, { rwa [ucs_succ, ih], }, } -/-- If a Lie module `M` contains a self-centralizing Lie submodule `N`, then all terms of the upper +/-- If a Lie module `M` contains a self-normalizing Lie submodule `N`, then all terms of the upper central series of `M` are contained in `N`. An important instance of this situation arises from a Cartan subalgebra `H ⊆ L` with the roles of `L`, `M`, `N` played by `H`, `L`, `H`, respectively. -/ -lemma ucs_le_of_centralizer_eq_self (h : N₁.centralizer = N₁) (k : ℕ) : +lemma ucs_le_of_normalizer_eq_self (h : N₁.normalizer = N₁) (k : ℕ) : (⊥ : lie_submodule R L M).ucs k ≤ N₁ := -by { rw ← ucs_eq_self_of_centralizer_eq_self h k, mono, simp, } +by { rw ← ucs_eq_self_of_normalizer_eq_self h k, mono, simp, } lemma lcs_add_le_iff (l k : ℕ) : N₁.lcs (l + k) ≤ N₂ ↔ N₁.lcs l ≤ N₂.ucs k := @@ -389,7 +389,7 @@ begin revert l, induction k with k ih, { simp, }, intros l, - rw [(by abel : l + (k + 1) = l + 1 + k), ih, ucs_succ, lcs_succ, top_lie_le_iff_le_centralizer], + rw [(by abel : l + (k + 1) = l + 1 + k), ih, ucs_succ, lcs_succ, top_lie_le_iff_le_normalizer], end lemma lcs_le_iff (k : ℕ) : diff --git a/src/algebra/lie/centralizer.lean b/src/algebra/lie/normalizer.lean similarity index 75% rename from src/algebra/lie/centralizer.lean rename to src/algebra/lie/normalizer.lean index eb28c3e542be6..e5a30f89fef4d 100644 --- a/src/algebra/lie/centralizer.lean +++ b/src/algebra/lie/normalizer.lean @@ -8,28 +8,28 @@ import algebra.lie.ideal_operations import algebra.lie.quotient /-! -# The centralizer of a Lie submodule and the normalizer of a Lie subalgebra. +# The normalizer of a Lie submodules and subalgebras. -Given a Lie module `M` over a Lie subalgebra `L`, the centralizer of a Lie submodule `N ⊆ M` is +Given a Lie module `M` over a Lie subalgebra `L`, the normalizer of a Lie submodule `N ⊆ M` is the Lie submodule with underlying set `{ m | ∀ (x : L), ⁅x, m⁆ ∈ N }`. -The lattice of Lie submodules thus has two natural operations, the centralizer: `N ↦ N.centralizer` +The lattice of Lie submodules thus has two natural operations, the normalizer: `N ↦ N.normalizer` and the ideal operation: `N ↦ ⁅⊤, N⁆`; these are adjoint, i.e., they form a Galois connection. This adjointness is the reason that we may define nilpotency in terms of either the upper or lower central series. Given a Lie subalgebra `H ⊆ L`, we may regard `H` as a Lie submodule of `L` over `H`, and thus -consider the centralizer. This turns out to be a Lie subalgebra and is known as the normalizer. +consider the normalizer. This turns out to be a Lie subalgebra. ## Main definitions - * `lie_submodule.centralizer` + * `lie_submodule.normalizer` * `lie_subalgebra.normalizer` - * `lie_submodule.gc_top_lie_centralizer` + * `lie_submodule.gc_top_lie_normalizer` ## Tags -lie algebra, centralizer, normalizer +lie algebra, normalizer -/ variables {R L M M' : Type*} @@ -41,53 +41,53 @@ namespace lie_submodule variables (N : lie_submodule R L M) {N₁ N₂ : lie_submodule R L M} -/-- The centralizer of a Lie submodule. -/ -def centralizer : lie_submodule R L M := +/-- The normalizer of a Lie submodule. -/ +def normalizer : lie_submodule R L M := { carrier := { m | ∀ (x : L), ⁅x, m⁆ ∈ N }, add_mem' := λ m₁ m₂ hm₁ hm₂ x, by { rw lie_add, exact N.add_mem' (hm₁ x) (hm₂ x), }, zero_mem' := λ x, by simp, smul_mem' := λ t m hm x, by { rw lie_smul, exact N.smul_mem' t (hm x), }, lie_mem := λ x m hm y, by { rw leibniz_lie, exact N.add_mem' (hm ⁅y, x⁆) (N.lie_mem (hm y)), } } -@[simp] lemma mem_centralizer (m : M) : - m ∈ N.centralizer ↔ ∀ (x : L), ⁅x, m⁆ ∈ N := +@[simp] lemma mem_normalizer (m : M) : + m ∈ N.normalizer ↔ ∀ (x : L), ⁅x, m⁆ ∈ N := iff.rfl -lemma le_centralizer : N ≤ N.centralizer := +lemma le_normalizer : N ≤ N.normalizer := begin intros m hm, - rw mem_centralizer, + rw mem_normalizer, exact λ x, N.lie_mem hm, end -lemma centralizer_inf : - (N₁ ⊓ N₂).centralizer = N₁.centralizer ⊓ N₂.centralizer := +lemma normalizer_inf : + (N₁ ⊓ N₂).normalizer = N₁.normalizer ⊓ N₂.normalizer := by { ext, simp [← forall_and_distrib], } -@[mono] lemma monotone_centalizer : - monotone (centralizer : lie_submodule R L M → lie_submodule R L M) := +@[mono] lemma monotone_normalizer : + monotone (normalizer : lie_submodule R L M → lie_submodule R L M) := begin intros N₁ N₂ h m hm, - rw mem_centralizer at hm ⊢, + rw mem_normalizer at hm ⊢, exact λ x, h (hm x), end -@[simp] lemma comap_centralizer (f : M' →ₗ⁅R,L⁆ M) : - N.centralizer.comap f = (N.comap f).centralizer := +@[simp] lemma comap_normalizer (f : M' →ₗ⁅R,L⁆ M) : + N.normalizer.comap f = (N.comap f).normalizer := by { ext, simp, } -lemma top_lie_le_iff_le_centralizer (N' : lie_submodule R L M) : - ⁅(⊤ : lie_ideal R L), N⁆ ≤ N' ↔ N ≤ N'.centralizer := +lemma top_lie_le_iff_le_normalizer (N' : lie_submodule R L M) : + ⁅(⊤ : lie_ideal R L), N⁆ ≤ N' ↔ N ≤ N'.normalizer := by { rw lie_le_iff, tauto, } -lemma gc_top_lie_centralizer : - galois_connection (λ N : lie_submodule R L M, ⁅(⊤ : lie_ideal R L), N⁆) centralizer := -top_lie_le_iff_le_centralizer +lemma gc_top_lie_normalizer : + galois_connection (λ N : lie_submodule R L M, ⁅(⊤ : lie_ideal R L), N⁆) normalizer := +top_lie_le_iff_le_normalizer variables (R L M) -lemma centralizer_bot_eq_max_triv_submodule : - (⊥ : lie_submodule R L M).centralizer = lie_module.max_triv_submodule R L M := +lemma normalizer_bot_eq_max_triv_submodule : + (⊥ : lie_submodule R L M).normalizer = lie_module.max_triv_submodule R L M := rfl end lie_submodule @@ -96,15 +96,15 @@ namespace lie_subalgebra variables (H : lie_subalgebra R L) -/-- Regarding a Lie subalgebra `H ⊆ L` as a module over itself, its centralizer is in fact a Lie -subalgebra. This is called the normalizer of the Lie subalgebra. -/ +/-- Regarding a Lie subalgebra `H ⊆ L` as a module over itself, its normalizer is in fact a Lie +subalgebra. -/ def normalizer : lie_subalgebra R L := { lie_mem' := λ y z hy hz x, begin rw [coe_bracket_of_module, mem_to_lie_submodule, leibniz_lie, ← lie_skew y, ← sub_eq_add_neg], exact H.sub_mem (hz ⟨_, hy x⟩) (hy ⟨_, hz x⟩), end, - .. H.to_lie_submodule.centralizer } + .. H.to_lie_submodule.normalizer } lemma mem_normalizer_iff' (x : L) : x ∈ H.normalizer ↔ ∀ (y : L), (y ∈ H) → ⁅y, x⁆ ∈ H := by { rw subtype.forall', refl, } @@ -116,10 +116,10 @@ begin rw [← lie_skew, neg_mem_iff], end -lemma le_normalizer : H ≤ H.normalizer := H.to_lie_submodule.le_centralizer +lemma le_normalizer : H ≤ H.normalizer := H.to_lie_submodule.le_normalizer -lemma coe_centralizer_eq_normalizer : - (H.to_lie_submodule.centralizer : submodule R L) = H.normalizer := +lemma coe_normalizer_eq_normalizer : + (H.to_lie_submodule.normalizer : submodule R L) = H.normalizer := rfl variables {H} diff --git a/src/algebra/lie/weights.lean b/src/algebra/lie/weights.lean index e75bad2d9d938..a2a68d870af65 100644 --- a/src/algebra/lie/weights.lean +++ b/src/algebra/lie/weights.lean @@ -462,7 +462,7 @@ begin refine le_antisymm _ (le_zero_root_subalgebra R L H), suffices : root_space H 0 ≤ H.to_lie_submodule, { exact λ x hx, this hx, }, obtain ⟨k, hk⟩ := (root_space H 0).is_nilpotent_iff_exists_self_le_ucs.mp (by apply_instance), - exact hk.trans (lie_submodule.ucs_le_of_centralizer_eq_self (by simp) k), + exact hk.trans (lie_submodule.ucs_le_of_normalizer_eq_self (by simp) k), end lemma zero_root_subalgebra_eq_iff_is_cartan [is_noetherian R L] : From b3607518eac6170b152bc7a540c16cef27cc17ac Mon Sep 17 00:00:00 2001 From: sgouezel Date: Wed, 1 Feb 2023 20:47:36 +0000 Subject: [PATCH 0407/1291] feat(analysis/calculus/bump_function_findim): construct good bump functions on finite-dimensional spaces (#18095) On any finite-dimensional real vector space, we construct a smooth family of smooth functions indexed by `R > 1` which are equal to `1` on the ball `B 0 1`, and with support exactly `B 0 R`. This is the main ingredient to construct nice bump functions. --- .../calculus/bump_function_findim.lean | 374 ++++++++++++++++++ src/analysis/convolution.lean | 24 ++ 2 files changed, 398 insertions(+) diff --git a/src/analysis/calculus/bump_function_findim.lean b/src/analysis/calculus/bump_function_findim.lean index bd39456517103..0ea049edb8028 100644 --- a/src/analysis/calculus/bump_function_findim.lean +++ b/src/analysis/calculus/bump_function_findim.lean @@ -7,6 +7,7 @@ import analysis.calculus.specific_functions import analysis.calculus.series import analysis.convolution import data.set.pointwise.support +import measure_theory.measure.haar_lebesgue /-! # Bump functions in finite-dimensional vector spaces @@ -194,3 +195,376 @@ begin end end + +section + +namespace exists_cont_diff_bump_base + +/-- An auxiliary function to construct partitions of unity on finite-dimensional real vector spaces. +It is the characteristic function of the closed unit ball. -/ +def φ : E → ℝ := (closed_ball (0 : E) 1).indicator (λ y, (1 : ℝ)) + +variables [normed_space ℝ E] [finite_dimensional ℝ E] + +section helper_definitions + +variable (E) +lemma u_exists : ∃ u : E → ℝ, cont_diff ℝ ⊤ u ∧ + (∀ x, u x ∈ Icc (0 : ℝ) 1) ∧ (support u = ball 0 1) ∧ (∀ x, u (-x) = u x) := +begin + have A : is_open (ball (0 : E) 1), from is_open_ball, + obtain ⟨f, f_support, f_smooth, f_range⟩ : + ∃ (f : E → ℝ), f.support = ball (0 : E) 1 ∧ cont_diff ℝ ⊤ f ∧ set.range f ⊆ set.Icc 0 1, + from A.exists_smooth_support_eq, + have B : ∀ x, f x ∈ Icc (0 : ℝ) 1 := λ x, f_range (mem_range_self x), + refine ⟨λ x, (f x + f (-x)) / 2, _, _, _, _⟩, + { exact (f_smooth.add (f_smooth.comp cont_diff_neg)).div_const }, + { assume x, + split, + { linarith [(B x).1, (B (-x)).1] }, + { linarith [(B x).2, (B (-x)).2] } }, + { refine support_eq_iff.2 ⟨λ x hx, _, λ x hx, _⟩, + { apply ne_of_gt, + have : 0 < f x, + { apply lt_of_le_of_ne (B x).1 (ne.symm _), + rwa ← f_support at hx }, + linarith [(B (-x)).1] }, + { have I1 : x ∉ support f, by rwa f_support, + have I2 : -x ∉ support f, + { rw f_support, + simp only at hx, + simpa using hx }, + simp only [mem_support, not_not] at I1 I2, + simp only [I1, I2, add_zero, zero_div] } }, + { assume x, simp only [add_comm, neg_neg] } +end + +variable {E} + +/-- An auxiliary function to construct partitions of unity on finite-dimensional real vector spaces, +which is smooth, symmetric, and with support equal to the unit ball. -/ +def u (x : E) : ℝ := classical.some (u_exists E) x + +variable (E) +lemma u_smooth : cont_diff ℝ ⊤ (u : E → ℝ) := (classical.some_spec (u_exists E)).1 + +lemma u_continuous : continuous (u : E → ℝ) := (u_smooth E).continuous + +lemma u_support : support (u : E → ℝ) = ball 0 1 := (classical.some_spec (u_exists E)).2.2.1 + +lemma u_compact_support : has_compact_support (u : E → ℝ) := +begin + rw [has_compact_support_def, u_support, closure_ball (0 : E) one_ne_zero], + exact is_compact_closed_ball _ _, +end +variable {E} + +lemma u_nonneg (x : E) : 0 ≤ u x := ((classical.some_spec (u_exists E)).2.1 x).1 + +lemma u_le_one (x : E) : u x ≤ 1 := ((classical.some_spec (u_exists E)).2.1 x).2 + +lemma u_neg (x : E) : u (-x) = u x := (classical.some_spec (u_exists E)).2.2.2 x + +variables [measurable_space E] [borel_space E] + +local notation `μ` := measure_theory.measure.add_haar + +variable (E) +lemma u_int_pos : 0 < ∫ (x : E), u x ∂μ := +begin + refine (integral_pos_iff_support_of_nonneg u_nonneg _).mpr _, + { exact (u_continuous E).integrable_of_has_compact_support (u_compact_support E) }, + { rw u_support, exact measure_ball_pos _ _ zero_lt_one } +end +variable {E} + +/-- An auxiliary function to construct partitions of unity on finite-dimensional real vector spaces, +which is smooth, symmetric, with support equal to the ball of radius `D` and integral `1`. -/ +def W (D : ℝ) (x : E) : ℝ := ((∫ (x : E), u x ∂μ) * |D|^(finrank ℝ E))⁻¹ • u (D⁻¹ • x) + +lemma W_def (D : ℝ) : + (W D : E → ℝ) = λ x, ((∫ (x : E), u x ∂μ) * |D|^(finrank ℝ E))⁻¹ • u (D⁻¹ • x) := +by { ext1 x, refl } + +lemma W_nonneg (D : ℝ) (x : E) : 0 ≤ W D x := +begin + apply mul_nonneg _ (u_nonneg _), + apply inv_nonneg.2, + apply mul_nonneg (u_int_pos E).le, + apply pow_nonneg (abs_nonneg D) +end + +lemma W_mul_φ_nonneg (D : ℝ) (x y : E) : 0 ≤ W D y * φ (x - y) := +mul_nonneg (W_nonneg D y) (indicator_nonneg (by simp only [zero_le_one, implies_true_iff]) _) + +variable (E) + +lemma W_integral {D : ℝ} (Dpos : 0 < D) : ∫ (x : E), W D x ∂μ = 1 := +begin + simp_rw [W, integral_smul], + rw [integral_comp_inv_smul_of_nonneg μ (u : E → ℝ) Dpos.le, + abs_of_nonneg Dpos.le, mul_comm], + field_simp [Dpos.ne', (u_int_pos E).ne'], +end + +lemma W_support {D : ℝ} (Dpos : 0 < D) : support (W D : E → ℝ) = ball 0 D := +begin + have B : D • ball (0 : E) 1 = ball 0 D, + by rw [smul_unit_ball Dpos.ne', real.norm_of_nonneg Dpos.le], + have C : D ^ finrank ℝ E ≠ 0, from pow_ne_zero _ Dpos.ne', + simp only [W_def, algebra.id.smul_eq_mul, support_mul, support_inv, univ_inter, + support_comp_inv_smul₀ Dpos.ne', u_support, B, support_const (u_int_pos E).ne', + support_const C, abs_of_nonneg Dpos.le], +end + +lemma W_compact_support {D : ℝ} (Dpos : 0 < D) : has_compact_support (W D : E → ℝ) := +begin + rw [has_compact_support_def, W_support E Dpos, closure_ball (0 : E) Dpos.ne'], + exact is_compact_closed_ball _ _, +end +variable {E} + +/-- An auxiliary function to construct partitions of unity on finite-dimensional real vector spaces. +It is the convolution between a smooth function of integral `1` supported in the ball of radius `D`, +with the indicator function of the closed unit ball. Therefore, it is smooth, equal to `1` on the +ball of radius `1 - D`, with support equal to the ball of radius `1 + D`. -/ +def Y (D : ℝ) : E → ℝ := W D ⋆[lsmul ℝ ℝ, μ] φ + +lemma Y_neg (D : ℝ) (x : E) : Y D (-x) = Y D x := +begin + apply convolution_neg_of_neg_eq, + { apply eventually_of_forall (λ x, _), + simp only [W_def, u_neg, smul_neg, algebra.id.smul_eq_mul, mul_eq_mul_left_iff, + eq_self_iff_true, true_or], }, + { apply eventually_of_forall (λ x, _), + simp only [φ, indicator, mem_closed_ball_zero_iff, norm_neg] }, +end + +lemma Y_eq_one_of_mem_closed_ball {D : ℝ} {x : E} (Dpos : 0 < D) + (hx : x ∈ closed_ball (0 : E) (1 - D)) : Y D x = 1 := +begin + change (W D ⋆[lsmul ℝ ℝ, μ] φ) x = 1, + have B : ∀ (y : E), y ∈ ball x D → φ y = 1, + { have C : ball x D ⊆ ball 0 1, + { apply ball_subset_ball', + simp only [mem_closed_ball] at hx, + linarith only [hx] }, + assume y hy, + simp only [φ, indicator, mem_closed_ball, ite_eq_left_iff, not_le, zero_ne_one], + assume h'y, + linarith only [mem_ball.1 (C hy), h'y] }, + have Bx : φ x = 1, from B _ (mem_ball_self Dpos), + have B' : ∀ y, y ∈ ball x D → φ y = φ x, by { rw Bx, exact B }, + rw convolution_eq_right' _ (le_of_eq (W_support E Dpos)) B', + simp only [lsmul_apply, algebra.id.smul_eq_mul, integral_mul_right, W_integral E Dpos, Bx, + one_mul], +end + +lemma Y_eq_zero_of_not_mem_ball {D : ℝ} {x : E} (Dpos : 0 < D) + (hx : x ∉ ball (0 : E) (1 + D)) : Y D x = 0 := +begin + change (W D ⋆[lsmul ℝ ℝ, μ] φ) x = 0, + have B : ∀ y, y ∈ ball x D → φ y = 0, + { assume y hy, + simp only [φ, indicator, mem_closed_ball_zero_iff, ite_eq_right_iff, one_ne_zero], + assume h'y, + have C : ball y D ⊆ ball 0 (1+D), + { apply ball_subset_ball', + rw ← dist_zero_right at h'y, + linarith only [h'y] }, + exact hx (C (mem_ball_comm.1 hy)) }, + have Bx : φ x = 0, from B _ (mem_ball_self Dpos), + have B' : ∀ y, y ∈ ball x D → φ y = φ x, by { rw Bx, exact B }, + rw convolution_eq_right' _ (le_of_eq (W_support E Dpos)) B', + simp only [lsmul_apply, algebra.id.smul_eq_mul, Bx, mul_zero, integral_const] +end + +lemma Y_nonneg (D : ℝ) (x : E) : 0 ≤ Y D x := +integral_nonneg (W_mul_φ_nonneg D x) + +lemma Y_le_one {D : ℝ} (x : E) (Dpos : 0 < D) : Y D x ≤ 1 := +begin + have A : (W D ⋆[lsmul ℝ ℝ, μ] φ) x ≤ (W D ⋆[lsmul ℝ ℝ, μ] 1) x, + { apply convolution_mono_right_of_nonneg _ (W_nonneg D) + (indicator_le_self' (λ x hx, zero_le_one)) (λ x, zero_le_one), + refine (has_compact_support.convolution_exists_left _ (W_compact_support E Dpos) _ + (locally_integrable_const (1 : ℝ)) x).integrable, + exact continuous_const.mul ((u_continuous E).comp (continuous_id.const_smul _)) }, + have B : (W D ⋆[lsmul ℝ ℝ, μ] (λ y, (1 : ℝ))) x = 1, + by simp only [convolution, continuous_linear_map.map_smul, mul_inv_rev, coe_smul', mul_one, + lsmul_apply, algebra.id.smul_eq_mul, integral_mul_left, W_integral E Dpos, pi.smul_apply], + exact A.trans (le_of_eq B) +end + +lemma Y_pos_of_mem_ball {D : ℝ} {x : E} (Dpos : 0 < D) (D_lt_one : D < 1) + (hx : x ∈ ball (0 : E) (1 + D)) : 0 < Y D x := +begin + simp only [mem_ball_zero_iff] at hx, + refine (integral_pos_iff_support_of_nonneg (W_mul_φ_nonneg D x) _).2 _, + { have F_comp : has_compact_support (W D), + from W_compact_support E Dpos, + have B : locally_integrable (φ : E → ℝ) μ, + from (locally_integrable_const _).indicator measurable_set_closed_ball, + have C : continuous (W D : E → ℝ), + from continuous_const.mul ((u_continuous E).comp (continuous_id.const_smul _)), + exact (has_compact_support.convolution_exists_left (lsmul ℝ ℝ : ℝ →L[ℝ] ℝ →L[ℝ] ℝ) + F_comp C B x).integrable }, + { set z := (D / (1 + D)) • x with hz, + have B : 0 < 1 + D, by linarith, + have C : ball z (D * (1 + D- ‖x‖) / (1 + D)) ⊆ support (λ (y : E), W D y * φ (x - y)), + { assume y hy, + simp only [support_mul, W_support E Dpos], + simp only [φ, mem_inter_iff, mem_support, ne.def, indicator_apply_eq_zero, + mem_closed_ball_zero_iff, one_ne_zero, not_forall, not_false_iff, exists_prop, and_true], + split, + { apply ball_subset_ball' _ hy, + simp only [z, norm_smul, abs_of_nonneg Dpos.le, abs_of_nonneg B.le, dist_zero_right, + real.norm_eq_abs, abs_div], + simp only [div_le_iff B] with field_simps, + ring_nf }, + { have ID : ‖D / (1 + D) - 1‖ = 1 / (1 + D), + { rw real.norm_of_nonpos, + { simp only [B.ne', ne.def, not_false_iff, mul_one, neg_sub, add_tsub_cancel_right] + with field_simps}, + { simp only [B.ne', ne.def, not_false_iff, mul_one] with field_simps, + apply div_nonpos_of_nonpos_of_nonneg _ B.le, + linarith only, } }, + rw ← mem_closed_ball_iff_norm', + apply closed_ball_subset_closed_ball' _ (ball_subset_closed_ball hy), + rw [← one_smul ℝ x, dist_eq_norm, hz, ← sub_smul, one_smul, norm_smul, ID], + simp only [-one_div, -mul_eq_zero, B.ne', div_le_iff B] with field_simps, + simp only [mem_ball_zero_iff] at hx, + nlinarith only [hx, D_lt_one] } }, + apply lt_of_lt_of_le _ (measure_mono C), + apply measure_ball_pos, + exact div_pos (mul_pos Dpos (by linarith only [hx])) B } +end + +variable (E) + +lemma Y_smooth : cont_diff_on ℝ ⊤ (uncurry Y) ((Ioo (0 : ℝ) 1) ×ˢ (univ : set E)) := +begin + have hs : is_open (Ioo (0 : ℝ) (1 : ℝ)), from is_open_Ioo, + have hk : is_compact (closed_ball (0 : E) 1), from proper_space.is_compact_closed_ball _ _, + refine cont_diff_on_convolution_left_with_param (lsmul ℝ ℝ) hs hk _ _ _, + { rintros p x hp hx, + simp only [W, mul_inv_rev, algebra.id.smul_eq_mul, mul_eq_zero, inv_eq_zero], + right, + contrapose! hx, + have : p⁻¹ • x ∈ support u, from mem_support.2 hx, + simp only [u_support, norm_smul, mem_ball_zero_iff, real.norm_eq_abs, abs_inv, + abs_of_nonneg hp.1.le, ← div_eq_inv_mul, div_lt_one hp.1] at this, + rw mem_closed_ball_zero_iff, + exact this.le.trans hp.2.le }, + { exact (locally_integrable_const _).indicator measurable_set_closed_ball }, + { apply cont_diff_on.mul, + { refine (cont_diff_on_const.mul _).inv + (λ x hx, ne_of_gt (mul_pos (u_int_pos E) (pow_pos (abs_pos_of_pos hx.1.1) _))), + apply cont_diff_on.pow, + simp_rw [← real.norm_eq_abs], + apply @cont_diff_on.norm ℝ, + { exact cont_diff_on_fst }, + { assume x hx, exact ne_of_gt hx.1.1 } }, + { apply (u_smooth E).comp_cont_diff_on, + exact cont_diff_on.smul (cont_diff_on_fst.inv (λ x hx, ne_of_gt hx.1.1)) cont_diff_on_snd } }, +end + +lemma Y_support {D : ℝ} (Dpos : 0 < D) (D_lt_one : D < 1) : + support (Y D : E → ℝ) = ball (0 : E) (1 + D) := +support_eq_iff.2 ⟨λ x hx, (Y_pos_of_mem_ball Dpos D_lt_one hx).ne', + λ x hx, Y_eq_zero_of_not_mem_ball Dpos hx⟩ + +variable {E} + +end helper_definitions + +/-- The base function from which one will construct a family of bump functions. One could +add more properties if they are useful and satisfied in the examples of inner product spaces +and finite dimensional vector spaces, notably derivative norm control in terms of `R - 1`. +TODO: move to the right file and use this to refactor bump functions in general. -/ +@[nolint has_nonempty_instance] +structure _root_.cont_diff_bump_base (E : Type*) [normed_add_comm_group E] [normed_space ℝ E] := +(to_fun : ℝ → E → ℝ) +(mem_Icc : ∀ (R : ℝ) (x : E), to_fun R x ∈ Icc (0 : ℝ) 1) +(symmetric : ∀ (R : ℝ) (x : E), to_fun R (-x) = to_fun R x) +(smooth : cont_diff_on ℝ ⊤ (uncurry to_fun) ((Ioi (1 : ℝ)) ×ˢ (univ : set E))) +(eq_one : ∀ (R : ℝ) (hR : 1 < R) (x : E) (hx : ‖x‖ ≤ 1), to_fun R x = 1) +(support : ∀ (R : ℝ) (hR : 1 < R), support (to_fun R) = metric.ball (0 : E) R) + +/-- A class registering that a real vector space admits bump functions. This will be instantiated +first for inner product spaces, and then for finite-dimensional normed spaces. +We use a specific class instead of `nonempty (cont_diff_bump_base E)` for performance reasons. -/ +class _root_.has_cont_diff_bump (E : Type*) [normed_add_comm_group E] [normed_space ℝ E] : Prop := +(out : nonempty (cont_diff_bump_base E)) + +@[priority 100] +instance {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] [finite_dimensional ℝ E] : + has_cont_diff_bump E := +begin + refine ⟨⟨_⟩⟩, + borelize E, + have IR : ∀ (R : ℝ), 1 < R → 0 < (R - 1) / (R + 1), + { assume R hR, apply div_pos; linarith }, + exact + { to_fun := λ R x, if 1 < R then Y ((R - 1) / (R + 1)) (((R + 1) / 2)⁻¹ • x) else 0, + mem_Icc := λ R x, begin + split_ifs, + { refine ⟨Y_nonneg _ _, Y_le_one _ (IR R h)⟩ }, + { simp only [pi.zero_apply, left_mem_Icc, zero_le_one] } + end, + symmetric := λ R x, begin + split_ifs, + { simp only [Y_neg, smul_neg] }, + { refl }, + end, + smooth := begin + suffices : cont_diff_on ℝ ⊤ + ((uncurry Y) ∘ (λ (p : ℝ × E), ((p.1 - 1) / (p.1 + 1), ((p.1 + 1)/2)⁻¹ • p.2))) + (Ioi 1 ×ˢ univ), + { apply this.congr, + rintros ⟨R, x⟩ ⟨(hR : 1 < R), hx⟩, + simp only [hR, uncurry_apply_pair, if_true, comp_app], }, + apply (Y_smooth E).comp, + { apply cont_diff_on.prod, + { refine (cont_diff_on_fst.sub cont_diff_on_const).div + (cont_diff_on_fst.add cont_diff_on_const) _, + rintros ⟨R, x⟩ ⟨(hR : 1 < R), hx⟩, + apply ne_of_gt, + dsimp only, + linarith, }, + { apply cont_diff_on.smul _ cont_diff_on_snd, + refine (cont_diff_on_fst.add cont_diff_on_const).div_const.inv _, + rintros ⟨R, x⟩ ⟨(hR : 1 < R), hx⟩, + apply ne_of_gt, + dsimp only, + linarith } }, + { rintros ⟨R, x⟩ ⟨(hR : 1 < R), hx⟩, + have A : 0 < (R - 1) / (R + 1), by { apply div_pos; linarith }, + have B : (R - 1) / (R + 1) < 1, by { apply (div_lt_one _ ).2; linarith }, + simp only [mem_preimage, prod_mk_mem_set_prod_eq, mem_Ioo, mem_univ, and_true, A, B] } + end, + eq_one := λ R hR x hx, begin + have A : 0 < R + 1, by linarith, + simp only [hR, if_true], + apply Y_eq_one_of_mem_closed_ball (IR R hR), + simp only [norm_smul, inv_div, mem_closed_ball_zero_iff, real.norm_eq_abs, abs_div, + abs_two, abs_of_nonneg A.le], + calc 2 / (R + 1) * ‖x‖ ≤ 2 / (R + 1) * 1 : + mul_le_mul_of_nonneg_left hx (div_nonneg zero_le_two A.le) + ... = 1 - (R - 1) / (R + 1) : by { field_simp [A.ne'], ring } + end, + support := λ R hR, begin + have A : 0 < (R + 1) / 2, by linarith, + have A' : 0 < R + 1, by linarith, + have C : (R - 1) / (R + 1) < 1, by { apply (div_lt_one _ ).2; linarith }, + simp only [hR, if_true, support_comp_inv_smul₀ A.ne', Y_support _ (IR R hR) C, + smul_ball A.ne', real.norm_of_nonneg A.le, smul_zero], + congr' 1, + field_simp [A'.ne'], + ring, + end }, +end + +end exists_cont_diff_bump_base + +end diff --git a/src/analysis/convolution.lean b/src/analysis/convolution.lean index 4cb07c26156f9..64b87646fa83f 100644 --- a/src/analysis/convolution.lean +++ b/src/analysis/convolution.lean @@ -482,6 +482,30 @@ lemma convolution_exists.add_distrib (hfg : convolution_exists f g L μ) (hfg' : convolution_exists f' g L μ) : (f + f') ⋆[L, μ] g = f ⋆[L, μ] g + f' ⋆[L, μ] g := by { ext, exact (hfg x).add_distrib (hfg' x) } +lemma convolution_mono_right {f g g' : G → ℝ} + (hfg : convolution_exists_at f g x (lsmul ℝ ℝ) μ) + (hfg' : convolution_exists_at f g' x (lsmul ℝ ℝ) μ) + (hf : ∀ x, 0 ≤ f x) (hg : ∀ x, g x ≤ g' x) : + (f ⋆[lsmul ℝ ℝ, μ] g) x ≤ (f ⋆[lsmul ℝ ℝ, μ] g') x := +begin + apply integral_mono hfg hfg', + simp only [lsmul_apply, algebra.id.smul_eq_mul], + assume t, + apply mul_le_mul_of_nonneg_left (hg _) (hf _), +end + +lemma convolution_mono_right_of_nonneg {f g g' : G → ℝ} + (hfg' : convolution_exists_at f g' x (lsmul ℝ ℝ) μ) + (hf : ∀ x, 0 ≤ f x) (hg : ∀ x, g x ≤ g' x) (hg' : ∀ x, 0 ≤ g' x) : + (f ⋆[lsmul ℝ ℝ, μ] g) x ≤ (f ⋆[lsmul ℝ ℝ, μ] g') x := +begin + by_cases H : convolution_exists_at f g x (lsmul ℝ ℝ) μ, + { exact convolution_mono_right H hfg' hf hg }, + have : (f ⋆[lsmul ℝ ℝ, μ] g) x = 0 := integral_undef H, + rw this, + exact integral_nonneg (λ y, mul_nonneg (hf y) (hg' (x - y))), +end + variables (L) lemma convolution_congr [has_measurable_add₂ G] [has_measurable_neg G] From d90e4e186f1d18e375dcd4e5b5f6364b01cb3e46 Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Wed, 1 Feb 2023 20:47:37 +0000 Subject: [PATCH 0408/1291] feat(analysis/special_functions/gaussian): Fourier transform of Gaussian (#18338) Use contour integrals around a rectangular contour to show that the Gaussian is its own Fourier transform. Co-authored-by: Meow Co-authored-by: sgouezel --- src/analysis/special_functions/gaussian.lean | 200 +++++++++++++++++++ 1 file changed, 200 insertions(+) diff --git a/src/analysis/special_functions/gaussian.lean b/src/analysis/special_functions/gaussian.lean index ca021f84a0bb9..2c8c165862218 100644 --- a/src/analysis/special_functions/gaussian.lean +++ b/src/analysis/special_functions/gaussian.lean @@ -3,9 +3,13 @@ Copyright (c) 2022 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ + import analysis.special_functions.gamma import analysis.special_functions.polar_coord import analysis.convex.complex +import analysis.normed.group.basic +import analysis.complex.cauchy_integral +import measure_theory.group.integration /-! # Gaussian integral @@ -16,6 +20,17 @@ We prove various versions of the formula for the Gaussian integral: `∫ x:ℝ, exp (-b * x^2) = (π / b) ^ (1 / 2)`. * `integral_gaussian_Ioi` and `integral_gaussian_complex_Ioi`: variants for integrals over `Ioi 0`. * `complex.Gamma_one_half_eq`: the formula `Γ (1 / 2) = √π`. + +We also prove, more generally, that the Fourier transform of the Gaussian +is another Gaussian: + +* `integral_cexp_neg_mul_sq_add_const`: for all complex `b` and `c` with `0 < re b` we have + `∫ (x : ℝ), exp (-b * (x + c) ^ 2) = (π / b) ^ (1 / 2)`. +* `fourier_transform_gaussian`: for all complex `b` and `t` with `0 < re b`, we have + `∫ x:ℝ, exp (I * t * x) * exp (-b * x^2) = (π / b) ^ (1 / 2) * exp (-t ^ 2 / (4 * b))`. +* `fourier_transform_gaussian_pi`: a variant with `b` and `t` scaled to give a more symmetric + statement, `∫ x:ℝ, exp (2 * π * I * t * x) * exp (-π * b * x^2) = + (1 / b) ^ (1 / 2) * exp (-π * (1 / b) * t ^ 2)`. -/ noncomputable theory @@ -334,3 +349,188 @@ begin { simpa only [one_div, of_real_inv, of_real_bit0] using Gamma_of_real (1 / 2)}, { rw [sqrt_eq_rpow, of_real_cpow pi_pos.le, of_real_div, of_real_bit0, of_real_one] } end + +namespace gaussian_fourier +/-! ## Fourier transform of the Gaussian integral +-/ +open interval_integral +open_locale real + +variables {b : ℂ} + +/-- The integral of the Gaussian function over the vertical edges of a rectangle +with vertices at `(±T, 0)` and `(±T, c)`. -/ +def vertical_integral (b : ℂ) (c T : ℝ) : ℂ := +∫ (y : ℝ) in 0..c, I * (cexp (-b * (T + y * I) ^ 2) - cexp (-b * (T - y * I) ^ 2)) + +/-- Explicit formula for the norm of the Gaussian function along the vertical +edges. -/ +lemma norm_cexp_neg_mul_sq_add_mul_I (b : ℂ) (c T : ℝ) : + ‖cexp (-b * (T + c * I) ^ 2)‖ = exp (-(b.re * T ^ 2 - 2 * b.im * c * T - b.re * c ^ 2)) := +begin + rw [complex.norm_eq_abs, complex.abs_exp, neg_mul, neg_re, ←re_add_im b], + simp only [sq, re_add_im, mul_re, mul_im, add_re, add_im, of_real_re, of_real_im, I_re, I_im], + ring_nf, +end + +lemma norm_cexp_neg_mul_sq_add_mul_I' (hb : b.re ≠ 0) (c T : ℝ) : + ‖cexp (-b * (T + c * I) ^ 2)‖ = + exp (-(b.re * (T - b.im * c / b.re) ^ 2 - c ^ 2 * (b.im ^ 2 / b.re + b.re))) := +begin + have : (b.re * T ^ 2 - 2 * b.im * c * T - b.re * c ^ 2) = + b.re * (T - b.im * c / b.re) ^ 2 - c ^ 2 * (b.im ^ 2 / b.re + b.re), + { field_simp, ring }, + rw [norm_cexp_neg_mul_sq_add_mul_I, this], +end + +lemma vertical_integral_norm_le (hb : 0 < b.re) (c : ℝ) {T : ℝ} (hT : 0 ≤ T) : + ‖vertical_integral b c T‖ + ≤ 2 * |c| * exp (-(b.re * T ^ 2 - 2 * |b.im| * |c| * T - b.re * c ^ 2)) := +begin + -- first get uniform bound for integrand + have vert_norm_bound : ∀ {T : ℝ}, 0 ≤ T → ∀ {c y : ℝ}, |y| ≤ |c| → + ‖cexp (-b * (T + y * I) ^ 2)‖ ≤ exp (-(b.re * T ^ 2 - 2 * |b.im| * |c| * T - b.re * c ^ 2)), + { intros T hT c y hy, + rw [norm_cexp_neg_mul_sq_add_mul_I b, exp_le_exp, neg_le_neg_iff], + refine sub_le_sub (sub_le_sub (le_refl _) (mul_le_mul_of_nonneg_right _ hT)) _, + { conv_lhs {rw mul_assoc}, conv_rhs {rw mul_assoc}, + refine mul_le_mul_of_nonneg_left ((le_abs_self _).trans _) zero_le_two, + rw abs_mul, + exact mul_le_mul_of_nonneg_left hy (abs_nonneg _), }, + { refine mul_le_mul_of_nonneg_left _ hb.le, + rwa sq_le_sq, } }, + -- now main proof + refine (interval_integral.norm_integral_le_of_norm_le_const _).trans _, + swap 3, + { rw sub_zero, + conv_lhs { rw mul_comm }, + conv_rhs { conv { congr, rw mul_comm }, rw mul_assoc } }, + { intros y hy, + have absy : |y| ≤ |c|, + { rcases le_or_lt 0 c, + { rw uIoc_of_le h at hy, + rw [abs_of_nonneg h, abs_of_pos hy.1], + exact hy.2, }, + { rw uIoc_of_lt h at hy, + rw [abs_of_neg h, abs_of_nonpos hy.2, neg_le_neg_iff], + exact hy.1.le } }, + rw [norm_mul, complex.norm_eq_abs, abs_I, one_mul, two_mul], + refine (norm_sub_le _ _).trans (add_le_add (vert_norm_bound hT absy) _), + rw ←abs_neg y at absy, + simpa only [neg_mul, of_real_neg] using vert_norm_bound hT absy }, +end + +lemma tendsto_vertical_integral (hb : 0 < b.re) (c : ℝ) : + tendsto (vertical_integral b c) at_top (𝓝 0) := +begin + -- complete proof using squeeze theorem: + rw tendsto_zero_iff_norm_tendsto_zero, + refine tendsto_of_tendsto_of_tendsto_of_le_of_le' tendsto_const_nhds _ + (eventually_of_forall (λ _, norm_nonneg _)) + ((eventually_ge_at_top (0:ℝ)).mp (eventually_of_forall + (λ T hT, vertical_integral_norm_le hb c hT))), + rw (by ring : 0 = 2 * |c| * 0), + refine (tendsto_exp_at_bot.comp (tendsto_neg_at_top_at_bot.comp _)).const_mul _ , + apply tendsto_at_top_add_const_right, + simp_rw [sq, ←mul_assoc, ←sub_mul], + refine tendsto.at_top_mul_at_top (tendsto_at_top_add_const_right _ _ _) tendsto_id, + exact (tendsto_const_mul_at_top_of_pos hb).mpr tendsto_id, +end + +lemma integrable_cexp_neg_mul_sq_add_real_mul_I (hb : 0 < b.re) (c : ℝ) : + integrable (λ (x : ℝ), cexp (-b * (x + c * I) ^ 2)) := +begin + refine ⟨(complex.continuous_exp.comp (continuous_const.mul ((continuous_of_real.add + continuous_const).pow 2))).ae_strongly_measurable, _⟩, + rw ←has_finite_integral_norm_iff, + simp_rw [norm_cexp_neg_mul_sq_add_mul_I' hb.ne', neg_sub _ (c ^ 2 * _), + sub_eq_add_neg _ (b.re * _), real.exp_add], + suffices : integrable (λ (x : ℝ), exp (-(b.re * x ^ 2))), + { exact (integrable.comp_sub_right this (b.im * c / b.re)).has_finite_integral.const_mul _, }, + simp_rw ←neg_mul, + apply integrable_exp_neg_mul_sq hb, +end + +lemma integral_cexp_neg_mul_sq_add_real_mul_I (hb : 0 < b.re) (c : ℝ) : + ∫ (x : ℝ), cexp (-b * (x + c * I) ^ 2) = (π / b) ^ (1 / 2 : ℂ) := +begin + refine tendsto_nhds_unique (interval_integral_tendsto_integral + (integrable_cexp_neg_mul_sq_add_real_mul_I hb c) tendsto_neg_at_top_at_bot tendsto_id) _, + set I₁ := (λ T, ∫ (x : ℝ) in -T..T, cexp (-b * (x + c * I) ^ 2)) with HI₁, + let I₂ := λ (T : ℝ), ∫ (x : ℝ) in -T..T, cexp (-b * x ^ 2), + let I₄ := λ (T : ℝ), ∫ (y : ℝ) in 0..c, cexp (-b * (T + y * I) ^ 2), + let I₅ := λ (T : ℝ), ∫ (y : ℝ) in 0..c, cexp (-b * (-T + y * I) ^ 2), + have C : ∀ (T : ℝ), I₂ T - I₁ T + I * I₄ T - I * I₅ T = 0, + { assume T, + have := integral_boundary_rect_eq_zero_of_differentiable_on + (λ z, cexp (-b * z ^ 2)) (-T) (T + c * I) + (by { refine differentiable.differentiable_on (differentiable.const_mul _ _).cexp, + exact differentiable_pow 2, }), + simpa only [neg_im, of_real_im, neg_zero, of_real_zero, zero_mul, add_zero, neg_re, of_real_re, + add_re, mul_re, I_re, mul_zero, I_im, tsub_zero, add_im, mul_im, mul_one, zero_add, + algebra.id.smul_eq_mul, of_real_neg] using this }, + simp_rw [id.def, ←HI₁], + have : I₁ = λ (T : ℝ), I₂ T + vertical_integral b c T, + { ext1 T, + specialize C T, + rw sub_eq_zero at C, + unfold vertical_integral, + rw [integral_const_mul, interval_integral.integral_sub], + { simp_rw (λ a b, by { rw sq, ring_nf } : ∀ (a b : ℂ), (a - b * I)^2 = (- a + b * I)^2), + change I₁ T = I₂ T + I * (I₄ T - I₅ T), + rw [mul_sub, ←C], + abel }, + all_goals { apply continuous.interval_integrable, continuity }, }, + rw [this, ←add_zero ((π / b : ℂ) ^ (1 / 2 : ℂ)), ←integral_gaussian_complex hb], + refine tendsto.add _ (tendsto_vertical_integral hb c), + exact interval_integral_tendsto_integral (integrable_cexp_neg_mul_sq hb) + tendsto_neg_at_top_at_bot tendsto_id, +end + +lemma _root_.integral_cexp_neg_mul_sq_add_const (hb : 0 < b.re) (c : ℂ) : + ∫ (x : ℝ), cexp (-b * (x + c) ^ 2) = (π / b) ^ (1 / 2 : ℂ) := +begin + rw ←re_add_im c, + simp_rw [←add_assoc, ←of_real_add], + rw integral_add_right_eq_self (λ(x : ℝ), cexp (-b * (↑x + ↑(c.im) * I) ^ 2)), + { apply integral_cexp_neg_mul_sq_add_real_mul_I hb }, + { apply_instance }, +end + +lemma _root_.fourier_transform_gaussian (hb : 0 < b.re) (t : ℂ) : + ∫ (x : ℝ), cexp (I * t * x) * cexp (-b * x ^ 2) = cexp (-t^2 / (4 * b)) * (π / b) ^ (1 / 2 : ℂ) := +begin + have : b ≠ 0, + { contrapose! hb, rw [hb, zero_re] }, + simp_rw [←complex.exp_add], + have : ∀ (x : ℂ), I * t * x + (-b * x ^ 2) = -t ^ 2 / (4 * b) + -b * (x + (-I * t / 2 / b)) ^ 2, + { intro x, + ring_nf SOP, + rw I_sq, + field_simp, ring }, + simp_rw [this, complex.exp_add, integral_mul_left, integral_cexp_neg_mul_sq_add_const hb] +end + +lemma _root_.fourier_transform_gaussian_pi (hb : 0 < b.re) (t : ℂ) : + ∫ (x : ℝ), cexp (2 * π * I * t * x) * cexp (-π * b * x ^ 2) = + 1 / b ^ (1 / 2 : ℂ) * cexp (-π * (1 / b) * t ^ 2) := +begin + have h1 : 0 < re (π * b) := by { rw of_real_mul_re, exact mul_pos pi_pos hb }, + have h2 : b ≠ 0 := by { contrapose! hb, rw [hb, zero_re], }, + convert _root_.fourier_transform_gaussian h1 (2 * π * t) using 1, + { congr' 1, + ext1 x, + congr' 2, + all_goals { ring } }, + { conv_lhs { rw mul_comm }, + congr' 2, + { field_simp [of_real_ne_zero.mpr pi_ne_zero], ring, }, + { rw [←div_div, div_self (of_real_ne_zero.mpr pi_ne_zero), cpow_def_of_ne_zero h2, + cpow_def_of_ne_zero (one_div_ne_zero h2), one_div, ←complex.exp_neg, ←neg_mul], + congr' 2, + rw [one_div, complex.log_inv], + rw [ne.def, arg_eq_pi_iff, not_and_distrib, not_lt], + exact or.inl hb.le } }, +end + +end gaussian_fourier From 4ef778c2d868d47bd1dc8b95039c305b77c443c9 Mon Sep 17 00:00:00 2001 From: leanprover-community-bot Date: Thu, 2 Feb 2023 05:09:37 +0000 Subject: [PATCH 0409/1291] chore(scripts): update nolints.txt (#18357) I am happy to remove some nolints for you! --- scripts/nolints.txt | 3 --- 1 file changed, 3 deletions(-) diff --git a/scripts/nolints.txt b/scripts/nolints.txt index 52ccde41118cc..a289f7daf787f 100644 --- a/scripts/nolints.txt +++ b/scripts/nolints.txt @@ -342,9 +342,6 @@ apply_nolint stream.is_bisimulation doc_blame -- group_theory/group_action/sub_mul_action.lean apply_nolint sub_mul_action.has_zero fails_quickly --- linear_algebra/affine_space/basis.lean -apply_nolint affine_basis.ext_elem fintype_finite - -- linear_algebra/affine_space/matrix.lean apply_nolint affine_basis.affine_independent_of_to_matrix_right_inv fintype_finite apply_nolint affine_basis.affine_span_eq_top_of_to_matrix_left_inv fintype_finite From f16e7a22e11fc09c71f25446ac1db23a24e8a0bd Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Thu, 2 Feb 2023 06:16:55 +0000 Subject: [PATCH 0410/1291] chore(*): add mathlib4 synchronization comments (#18355) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.big_operators.fin` * `data.finset.finsupp` * `data.finset.functor` * `data.finset.pointwise` * `data.finsupp.basic` * `data.finsupp.big_operators` * `data.finsupp.pwo` * `data.multiset.locally_finite` * `data.set.pointwise.interval` * `group_theory.subgroup.pointwise` * `group_theory.submonoid.pointwise` * `order.compactly_generated` * `order.minimal` * `order.upper_lower.basic` * `order.well_founded_set` * `ring_theory.subsemiring.pointwise` * `topology.bases` * `topology.subset_properties` --- src/algebra/big_operators/fin.lean | 3 +++ src/data/finset/finsupp.lean | 3 +++ src/data/finset/functor.lean | 3 +++ src/data/finset/pointwise.lean | 3 +++ src/data/finsupp/basic.lean | 3 +++ src/data/finsupp/big_operators.lean | 3 +++ src/data/finsupp/pwo.lean | 3 +++ src/data/multiset/locally_finite.lean | 3 +++ src/data/set/pointwise/interval.lean | 3 +++ src/group_theory/subgroup/pointwise.lean | 3 +++ src/group_theory/submonoid/pointwise.lean | 3 +++ src/order/compactly_generated.lean | 3 +++ src/order/minimal.lean | 3 +++ src/order/upper_lower/basic.lean | 3 +++ src/order/well_founded_set.lean | 3 +++ src/ring_theory/subsemiring/pointwise.lean | 3 +++ src/topology/bases.lean | 3 +++ src/topology/subset_properties.lean | 3 +++ 18 files changed, 54 insertions(+) diff --git a/src/algebra/big_operators/fin.lean b/src/algebra/big_operators/fin.lean index 20af1963a7386..a6d64ed604a40 100644 --- a/src/algebra/big_operators/fin.lean +++ b/src/algebra/big_operators/fin.lean @@ -11,6 +11,9 @@ import logic.equiv.fin /-! # Big operators and `fin` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Some results about products and sums over the type `fin`. The most important results are the induction formulas `fin.prod_univ_cast_succ` diff --git a/src/data/finset/finsupp.lean b/src/data/finset/finsupp.lean index 79b9b47bdf144..9bed84e1c0669 100644 --- a/src/data/finset/finsupp.lean +++ b/src/data/finset/finsupp.lean @@ -11,6 +11,9 @@ import data.fintype.big_operators /-! # Finitely supported product of finsets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the finitely supported product of finsets as a `finset (ι →₀ α)`. ## Main declarations diff --git a/src/data/finset/functor.lean b/src/data/finset/functor.lean index d8697c8891dfa..32dd58644474c 100644 --- a/src/data/finset/functor.lean +++ b/src/data/finset/functor.lean @@ -10,6 +10,9 @@ import data.multiset.functor /-! # Functoriality of `finset` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the functor structure of `finset`. ## TODO diff --git a/src/data/finset/pointwise.lean b/src/data/finset/pointwise.lean index b73271640aefe..1e750480d0b8e 100644 --- a/src/data/finset/pointwise.lean +++ b/src/data/finset/pointwise.lean @@ -11,6 +11,9 @@ import data.set.pointwise.list_of_fn /-! # Pointwise operations of finsets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines pointwise algebraic operations on finsets. ## Main declarations diff --git a/src/data/finsupp/basic.lean b/src/data/finsupp/basic.lean index 516f5255164bd..4b1e6bda95cf5 100644 --- a/src/data/finsupp/basic.lean +++ b/src/data/finsupp/basic.lean @@ -12,6 +12,9 @@ import data.rat.big_operators /-! # Miscellaneous definitions, lemmas, and constructions using finsupp +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main declarations * `finsupp.graph`: the finset of input and output pairs with non-zero outputs. diff --git a/src/data/finsupp/big_operators.lean b/src/data/finsupp/big_operators.lean index f91e7b90d5c8b..9f0051a7f7e10 100644 --- a/src/data/finsupp/big_operators.lean +++ b/src/data/finsupp/big_operators.lean @@ -10,6 +10,9 @@ import data.finset.pairwise /-! # Sums of collections of finsupp, and their support + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. This file provides results about the `finsupp.support` of sums of collections of `finsupp`, including sums of `list`, `multiset`, and `finset`. diff --git a/src/data/finsupp/pwo.lean b/src/data/finsupp/pwo.lean index 453b5157c147d..572dbef6e3bc1 100644 --- a/src/data/finsupp/pwo.lean +++ b/src/data/finsupp/pwo.lean @@ -9,6 +9,9 @@ import order.well_founded_set /-! # Partial well ordering on finsupps +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains the fact that finitely supported functions from a fintype are partially well ordered when the codomain is a linear order that is well ordered. It is in a separate file for now so as to not add imports to the file `order.well_founded_set`. diff --git a/src/data/multiset/locally_finite.lean b/src/data/multiset/locally_finite.lean index 988094e14a917..78d3b6f2d0a1d 100644 --- a/src/data/multiset/locally_finite.lean +++ b/src/data/multiset/locally_finite.lean @@ -8,6 +8,9 @@ import data.finset.locally_finite /-! # Intervals as multisets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file provides basic results about all the `multiset.Ixx`, which are defined in `order.locally_finite`. diff --git a/src/data/set/pointwise/interval.lean b/src/data/set/pointwise/interval.lean index d62629b46dcdd..b463eabdd2e8d 100644 --- a/src/data/set/pointwise/interval.lean +++ b/src/data/set/pointwise/interval.lean @@ -12,6 +12,9 @@ import algebra.order.group.min_max /-! # (Pre)images of intervals +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove a bunch of trivial lemmas like “if we add `a` to all points of `[b, c]`, then we get `[a + b, a + c]`”. For the functions `x ↦ x ± a`, `x ↦ a ± x`, and `x ↦ -x` we prove lemmas about preimages and images of all intervals. We also prove a few lemmas about images under diff --git a/src/group_theory/subgroup/pointwise.lean b/src/group_theory/subgroup/pointwise.lean index 3a565dfb0313f..91cb1a73b5073 100644 --- a/src/group_theory/subgroup/pointwise.lean +++ b/src/group_theory/subgroup/pointwise.lean @@ -9,6 +9,9 @@ import group_theory.group_action.conj_act /-! # Pointwise instances on `subgroup` and `add_subgroup`s +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file provides the actions * `subgroup.pointwise_mul_action` diff --git a/src/group_theory/submonoid/pointwise.lean b/src/group_theory/submonoid/pointwise.lean index c9a53990dacba..599da656d223b 100644 --- a/src/group_theory/submonoid/pointwise.lean +++ b/src/group_theory/submonoid/pointwise.lean @@ -9,6 +9,9 @@ import order.well_founded_set /-! # Pointwise instances on `submonoid`s and `add_submonoid`s +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file provides: * `submonoid.has_inv` diff --git a/src/order/compactly_generated.lean b/src/order/compactly_generated.lean index 79ad5b3cdeb25..6263ce4ff2347 100644 --- a/src/order/compactly_generated.lean +++ b/src/order/compactly_generated.lean @@ -16,6 +16,9 @@ import tactic.tfae /-! # Compactness properties for complete lattices +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + For complete lattices, there are numerous equivalent ways to express the fact that the relation `>` is well-founded. In this file we define three especially-useful characterisations and provide proofs that they are indeed equivalent to well-foundedness. diff --git a/src/order/minimal.lean b/src/order/minimal.lean index f898c52d85c3e..0010e217b008a 100644 --- a/src/order/minimal.lean +++ b/src/order/minimal.lean @@ -9,6 +9,9 @@ import order.upper_lower.basic /-! # Minimal/maximal elements of a set +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines minimal and maximal of a set with respect to an arbitrary relation. ## Main declarations diff --git a/src/order/upper_lower/basic.lean b/src/order/upper_lower/basic.lean index 262f5b7c9283e..bfc501061dd55 100644 --- a/src/order/upper_lower/basic.lean +++ b/src/order/upper_lower/basic.lean @@ -10,6 +10,9 @@ import data.set.intervals.order_iso /-! # Up-sets and down-sets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines upper and lower sets in an order. ## Main declarations diff --git a/src/order/well_founded_set.lean b/src/order/well_founded_set.lean index f726237ba0da4..37a675eb498cb 100644 --- a/src/order/well_founded_set.lean +++ b/src/order/well_founded_set.lean @@ -11,6 +11,9 @@ import tactic.tfae /-! # Well-founded sets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A well-founded subset of an ordered type is one on which the relation `<` is well-founded. ## Main Definitions diff --git a/src/ring_theory/subsemiring/pointwise.lean b/src/ring_theory/subsemiring/pointwise.lean index 4ac6690cf4c33..602d99c4ccd9b 100644 --- a/src/ring_theory/subsemiring/pointwise.lean +++ b/src/ring_theory/subsemiring/pointwise.lean @@ -10,6 +10,9 @@ import data.set.pointwise.basic /-! # Pointwise instances on `subsemiring`s +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file provides the action `subsemiring.pointwise_mul_action` which matches the action of `mul_action_set`. diff --git a/src/topology/bases.lean b/src/topology/bases.lean index cebad1c82cd8a..81f4a211ad817 100644 --- a/src/topology/bases.lean +++ b/src/topology/bases.lean @@ -9,6 +9,9 @@ import topology.continuous_on /-! # Bases of topologies. Countability axioms. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A topological basis on a topological space `t` is a collection of sets, such that all open sets can be generated as unions of these sets, without the need to take finite intersections of them. This file introduces a framework for dealing with these collections, diff --git a/src/topology/subset_properties.lean b/src/topology/subset_properties.lean index 0d23021b619db..d7200e25c410c 100644 --- a/src/topology/subset_properties.lean +++ b/src/topology/subset_properties.lean @@ -14,6 +14,9 @@ import order.minimal /-! # Properties of subsets of topological spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define various properties of subsets of a topological space, and some classes on topological spaces. From aebd342823838ca946af454f5b3194db77a52e65 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Mantas=20Bak=C5=A1ys?= Date: Thu, 2 Feb 2023 08:56:03 +0000 Subject: [PATCH 0411/1291] feat(algebra/order/chebyshev): Chebyshev's Sum Inequality (#13187) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Prove Chebyshev's sum inequality as a corollary to the rearrangement inequality. Co-authored-by: YaelDillies Co-authored-by: Mantas Bakšys <39908973+MantasBaksys@users.noreply.github.com> Co-authored-by: Yaël Dillies --- src/algebra/order/chebyshev.lean | 130 +++++++++++++++++++++++++++++++ 1 file changed, 130 insertions(+) create mode 100644 src/algebra/order/chebyshev.lean diff --git a/src/algebra/order/chebyshev.lean b/src/algebra/order/chebyshev.lean new file mode 100644 index 0000000000000..c825cf8a4f8b4 --- /dev/null +++ b/src/algebra/order/chebyshev.lean @@ -0,0 +1,130 @@ +/- +Copyright (c) 2023 Mantas Bakšys, Yaël Dillies. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Mantas Bakšys, Yaël Dillies +-/ +import algebra.big_operators.order +import algebra.order.rearrangement +import group_theory.perm.cycle.basic + +/-! +# Chebyshev's sum inequality + +This file proves the Chebyshev sum inequality. + +Chebyshev's inequality states `(∑ i in s, f i) * (∑ i in s, g i) ≤ s.card * ∑ i in s, f i * g i` +when `f g : ι → α` monovary, and the reverse inequality when `f` and `g` antivary. + + +## Main declarations + +* `monovary_on.sum_mul_sum_le_card_mul_sum`: Chebyshev's inequality. +* `antivary_on.card_mul_sum_le_sum_mul_sum`: Chebyshev's inequality, dual version. +* `sq_sum_le_card_mul_sum_sq`: Special case of Chebyshev's inequality when `f = g`. + +## Implementation notes + +In fact, we don't need much compatibility between the addition and multiplication of `α`, so we can +actually decouple them by replacing multiplication with scalar multiplication and making `f` and `g` +land in different types. +As a bonus, this makes the dual statement trivial. The multiplication versions are provided for +convenience. + +The case for `monotone`/`antitone` pairs of functions over a `linear_order` is not deduced in this +file because it is easily deducible from the `monovary` API. +-/ + +open equiv equiv.perm finset function order_dual +open_locale big_operators + +variables {ι α β : Type*} + +/-! ### Scalar multiplication versions -/ + +section smul +variables [linear_ordered_ring α] [linear_ordered_add_comm_group β] [module α β] + [ordered_smul α β] {s : finset ι} {σ : perm ι} {f : ι → α} {g : ι → β} + +/-- **Chebyshev's Sum Inequality**: When `f` and `g` monovary together (eg they are both +monotone/antitone), the scalar product of their sum is less than the size of the set times their +scalar product. -/ +lemma monovary_on.sum_smul_sum_le_card_smul_sum (hfg : monovary_on f g s) : + (∑ i in s, f i) • ∑ i in s, g i ≤ s.card • ∑ i in s, f i • g i := +begin + classical, + obtain ⟨σ, hσ, hs⟩ := s.countable_to_set.exists_cycle_on, + rw [←card_range s.card, sum_smul_sum_eq_sum_perm hσ], + exact sum_le_card_nsmul _ _ _ (λ n _, hfg.sum_smul_comp_perm_le_sum_smul $ λ x hx, hs $ λ h, hx $ + is_fixed_pt.perm_pow h _), +end + +/-- **Chebyshev's Sum Inequality**: When `f` and `g` antivary together (eg one is monotone, the +other is antitone), the scalar product of their sum is less than the size of the set times their +scalar product. -/ +lemma antivary_on.card_smul_sum_le_sum_smul_sum (hfg : antivary_on f g s) : + s.card • ∑ i in s, f i • g i ≤ (∑ i in s, f i) • ∑ i in s, g i := +by convert hfg.dual_right.sum_smul_sum_le_card_smul_sum + +variables [fintype ι] + +/-- **Chebyshev's Sum Inequality**: When `f` and `g` monovary together (eg they are both +monotone/antitone), the scalar product of their sum is less than the size of the set times their +scalar product. -/ +lemma monovary.sum_smul_sum_le_card_smul_sum (hfg : monovary f g) : + (∑ i, f i) • ∑ i, g i ≤ fintype.card ι • ∑ i, f i • g i := +(hfg.monovary_on _).sum_smul_sum_le_card_smul_sum + +/-- **Chebyshev's Sum Inequality**: When `f` and `g` antivary together (eg one is monotone, the +other is antitone), the scalar product of their sum is less than the size of the set times their +scalar product. -/ +lemma antivary.card_smul_sum_le_sum_smul_sum (hfg : antivary f g) : + fintype.card ι • ∑ i, f i • g i ≤ (∑ i, f i) • ∑ i, g i := +by convert (hfg.dual_right.monovary_on _).sum_smul_sum_le_card_smul_sum + +end smul + +/-! +### Multiplication versions + +Special cases of the above when scalar multiplication is actually multiplication. +-/ + +section mul +variables [linear_ordered_ring α] {s : finset ι} {σ : perm ι} {f g : ι → α} + +/-- **Chebyshev's Sum Inequality**: When `f` and `g` monovary together (eg they are both +monotone/antitone), the product of their sum is less than the size of the set times their scalar +product. -/ +lemma monovary_on.sum_mul_sum_le_card_mul_sum (hfg : monovary_on f g s) : + (∑ i in s, f i) * (∑ i in s, g i) ≤ s.card * ∑ i in s, f i * g i := +by { rw ←nsmul_eq_mul, exact hfg.sum_smul_sum_le_card_smul_sum } + +/-- **Chebyshev's Sum Inequality**: When `f` and `g` antivary together (eg one is monotone, the +other is antitone), the product of their sum is greater than the size of the set times their scalar +product. -/ +lemma antivary_on.card_mul_sum_le_sum_mul_sum (hfg : antivary_on f g s) : + (s.card : α) * ∑ i in s, f i * g i ≤ (∑ i in s, f i) * (∑ i in s, g i) := +by { rw ←nsmul_eq_mul, exact hfg.card_smul_sum_le_sum_smul_sum } + +/-- Special case of **Chebyshev's Sum Inequality** or the **Cauchy-Schwarz Inequality**: The square +of the sum is less than the size of the set times the sum of the squares. -/ +lemma sq_sum_le_card_mul_sum_sq : (∑ i in s, f i)^2 ≤ s.card * ∑ i in s, f i ^ 2 := +by { simp_rw sq, exact (monovary_on_self _ _).sum_mul_sum_le_card_mul_sum } + +variables [fintype ι] + +/-- **Chebyshev's Sum Inequality**: When `f` and `g` monovary together (eg they are both +monotone/antitone), the product of their sum is less than the size of the set times their scalar +product. -/ +lemma monovary.sum_mul_sum_le_card_mul_sum (hfg : monovary f g) : + (∑ i, f i) * (∑ i, g i) ≤ fintype.card ι * ∑ i, f i * g i := +(hfg.monovary_on _).sum_mul_sum_le_card_mul_sum + +/-- **Chebyshev's Sum Inequality**: When `f` and `g` antivary together (eg one is monotone, the +other is antitone), the product of their sum is less than the size of the set times their scalar +product. -/ +lemma antivary.card_mul_sum_le_sum_mul_sum (hfg : antivary f g) : + (fintype.card ι : α) * ∑ i, f i * g i ≤ (∑ i, f i) * (∑ i, g i) := +(hfg.antivary_on _).card_mul_sum_le_sum_mul_sum + +end mul From 2e59a6de168f95d16b16d217b808a36290398c0a Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Thu, 2 Feb 2023 10:32:45 +0000 Subject: [PATCH 0412/1291] =?UTF-8?q?feat(ring=5Ftheory/localization):=20r?= =?UTF-8?q?emove=20extraneous=20`algebra=5Fmap=5Fsubmonoid=20S=20M=20?= =?UTF-8?q?=E2=89=A4=20S=E2=81=B0`=20conditions=20(#18358)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- src/ring_theory/ideal/norm.lean | 15 +++--- src/ring_theory/localization/module.lean | 58 +++++++++++++----------- src/ring_theory/localization/norm.lean | 8 ++-- 3 files changed, 43 insertions(+), 38 deletions(-) diff --git a/src/ring_theory/ideal/norm.lean b/src/ring_theory/ideal/norm.lean index deb0ebf48daf5..9b7e4ad6aa35b 100644 --- a/src/ring_theory/ideal/norm.lean +++ b/src/ring_theory/ideal/norm.lean @@ -456,11 +456,10 @@ lemma span_norm_mono {I J : ideal S} (h : I ≤ J) : span_norm R I ≤ span_norm ideal.span_mono (set.monotone_image h) lemma span_norm_localization (I : ideal S) [module.finite R S] [module.free R S] - {M : submonoid R} {Rₘ Sₘ : Type*} + (M : submonoid R) {Rₘ Sₘ : Type*} [comm_ring Rₘ] [algebra R Rₘ] [comm_ring Sₘ] [algebra S Sₘ] [algebra Rₘ Sₘ] [algebra R Sₘ] [is_scalar_tower R Rₘ Sₘ] [is_scalar_tower R S Sₘ] - [is_localization M Rₘ] [is_localization (algebra.algebra_map_submonoid S M) Sₘ] - (hM : algebra.algebra_map_submonoid S M ≤ S⁰) : + [is_localization M Rₘ] [is_localization (algebra.algebra_map_submonoid S M) Sₘ] : span_norm Rₘ (I.map (algebra_map S Sₘ)) = (span_norm R I).map (algebra_map R Rₘ) := begin casesI h : subsingleton_or_nontrivial R, @@ -482,13 +481,13 @@ begin apply_fun algebra.norm Rₘ at has, rwa [_root_.map_mul, ← is_scalar_tower.algebra_map_apply, is_scalar_tower.algebra_map_apply R Rₘ, - algebra.norm_algebra_map_of_basis (b.localization_localization Rₘ M Sₘ hM), - algebra.norm_localization R a hM] at has, - all_goals { apply_instance} }, + algebra.norm_algebra_map_of_basis (b.localization_localization Rₘ M Sₘ), + algebra.norm_localization R M a] at has, + all_goals { apply_instance } }, { intros a ha, - rw [set.mem_preimage, function.comp_app, ← algebra.norm_localization R a hM], + rw [set.mem_preimage, function.comp_app, ← algebra.norm_localization R M a], exact subset_span (set.mem_image_of_mem _ (mem_map_of_mem _ ha)), - all_goals { apply_instance} }, + all_goals { apply_instance } }, end end ideal diff --git a/src/ring_theory/localization/module.lean b/src/ring_theory/localization/module.lean index b83bb0035a09f..e9ceea7284331 100644 --- a/src/ring_theory/localization/module.lean +++ b/src/ring_theory/localization/module.lean @@ -40,17 +40,14 @@ lemma linear_independent.localization {ι : Type*} {b : ι → M} (hli : linear_ begin rw linear_independent_iff' at ⊢ hli, intros s g hg i hi, - choose a g' hg' using is_localization.exist_integer_multiples S s g, - letI := λ i, classical.prop_decidable (i ∈ s), - specialize hli s (λ i, if hi : i ∈ s then g' i hi else 0) _ i hi, + choose! a g' hg' using is_localization.exist_integer_multiples S s g, + specialize hli s g' _ i hi, { rw [← @smul_zero _ M _ _ (a : R), ← hg, finset.smul_sum], refine finset.sum_congr rfl (λ i hi, _), - dsimp only, - rw [dif_pos hi, ← is_scalar_tower.algebra_map_smul Rₛ, hg' i hi, smul_assoc], + rw [← is_scalar_tower.algebra_map_smul Rₛ, hg' i hi, smul_assoc], apply_instance }, refine ((is_localization.map_units Rₛ a).mul_right_eq_zero).mp _, - rw [← algebra.smul_def, ← map_zero (algebra_map R Rₛ), ← hli], - simp [hi, hg'] + rw [← algebra.smul_def, ← map_zero (algebra_map R Rₛ), ← hli, hg' i hi], end end add_comm_monoid @@ -64,14 +61,26 @@ include hA open submodule -lemma linear_independent.localization_localization {ι : Type*} - {v : ι → A} (hv : linear_independent R v) (hS : algebra.algebra_map_submonoid A S ≤ A⁰) : +lemma linear_independent.localization_localization + {ι : Type*} {v : ι → A} (hv : linear_independent R v) : linear_independent Rₛ (algebra_map A Aₛ ∘ v) := begin - refine (hv.map' ((algebra.linear_map A Aₛ).restrict_scalars R) _).localization Rₛ S, - rw [linear_map.ker_restrict_scalars, restrict_scalars_eq_bot_iff, linear_map.ker_eq_bot, - algebra.coe_linear_map], - exact is_localization.injective Aₛ hS + rw linear_independent_iff' at ⊢ hv, + intros s g hg i hi, + choose! a g' hg' using is_localization.exist_integer_multiples S s g, + have h0 : algebra_map A Aₛ (∑ i in s, g' i • v i) = 0, + { apply_fun ((•) (a : R)) at hg, + rw [smul_zero, finset.smul_sum] at hg, + rw [map_sum, ← hg], + refine finset.sum_congr rfl (λ i hi, _), + rw [← smul_assoc, ← hg' i hi, algebra.smul_def, map_mul, + ← is_scalar_tower.algebra_map_apply, ← algebra.smul_def, algebra_map_smul] }, + obtain ⟨⟨_, r, hrS, rfl⟩, (hr : algebra_map R A r * _ = 0)⟩ := + (is_localization.map_eq_zero_iff (algebra.algebra_map_submonoid A S) _ _).1 h0, + simp_rw [finset.mul_sum, ← algebra.smul_def, smul_smul] at hr, + specialize hv s _ hr i hi, + rw [← (is_localization.map_units Rₛ a).mul_right_eq_zero, ← algebra.smul_def, ← hg' i hi], + exact (is_localization.map_eq_zero_iff S _ _).2 ⟨⟨r, hrS⟩, hv⟩, end lemma span_eq_top.localization_localization {v : set A} (hv : span R v = ⊤) : @@ -94,31 +103,28 @@ end A suitable instance for `[algebra A Aₛ]` is `localization_algebra`. -/ -noncomputable def basis.localization_localization {ι : Type*} (b : basis ι R A) - (hS : algebra.algebra_map_submonoid A S ≤ A⁰) : - basis ι Rₛ Aₛ := +noncomputable def basis.localization_localization {ι : Type*} (b : basis ι R A) : basis ι Rₛ Aₛ := basis.mk - (b.linear_independent.localization_localization _ S _ hS) + (b.linear_independent.localization_localization _ S _) (by { rw [set.range_comp, span_eq_top.localization_localization Rₛ S Aₛ b.span_eq], exact le_rfl }) -@[simp] lemma basis.localization_localization_apply {ι : Type*} (b : basis ι R A) - (hS : algebra.algebra_map_submonoid A S ≤ A⁰) (i) : - b.localization_localization Rₛ S Aₛ hS i = algebra_map A Aₛ (b i) := +@[simp] lemma basis.localization_localization_apply {ι : Type*} (b : basis ι R A) (i) : + b.localization_localization Rₛ S Aₛ i = algebra_map A Aₛ (b i) := basis.mk_apply _ _ _ -@[simp] lemma basis.localization_localization_repr_algebra_map {ι : Type*} (b : basis ι R A) - (hS : algebra.algebra_map_submonoid A S ≤ A⁰) (x i) : - (b.localization_localization Rₛ S Aₛ hS).repr (algebra_map A Aₛ x) i = +@[simp] lemma basis.localization_localization_repr_algebra_map + {ι : Type*} (b : basis ι R A) (x i) : + (b.localization_localization Rₛ S Aₛ).repr (algebra_map A Aₛ x) i = algebra_map R Rₛ (b.repr x i) := -calc (b.localization_localization Rₛ S Aₛ hS).repr (algebra_map A Aₛ x) i - = (b.localization_localization Rₛ S Aₛ hS).repr +calc (b.localization_localization Rₛ S Aₛ).repr (algebra_map A Aₛ x) i + = (b.localization_localization Rₛ S Aₛ).repr ((b.repr x).sum (λ j c, algebra_map R Rₛ c • algebra_map A Aₛ (b j))) i : by simp_rw [is_scalar_tower.algebra_map_smul, algebra.smul_def, is_scalar_tower.algebra_map_apply R A Aₛ, ← _root_.map_mul, ← map_finsupp_sum, ← algebra.smul_def, ← finsupp.total_apply, basis.total_repr] ... = (b.repr x).sum (λ j c, algebra_map R Rₛ c • finsupp.single j 1 i) : - by simp_rw [← b.localization_localization_apply Rₛ S Aₛ hS, map_finsupp_sum, + by simp_rw [← b.localization_localization_apply Rₛ S Aₛ, map_finsupp_sum, linear_equiv.map_smul, basis.repr_self, finsupp.sum_apply, finsupp.smul_apply] ... = _ : finset.sum_eq_single i (λ j _ hj, by simp [hj]) diff --git a/src/ring_theory/localization/norm.lean b/src/ring_theory/localization/norm.lean index 0e3b7a7d28a36..5eba4dbb697ea 100644 --- a/src/ring_theory/localization/norm.lean +++ b/src/ring_theory/localization/norm.lean @@ -29,15 +29,15 @@ open_locale non_zero_divisors variables (R : Type*) {S : Type*} [comm_ring R] [comm_ring S] [algebra R S] variables {Rₘ Sₘ : Type*} [comm_ring Rₘ] [algebra R Rₘ] [comm_ring Sₘ] [algebra S Sₘ] -variables {M : submonoid R} +variables (M : submonoid R) variables [is_localization M Rₘ] [is_localization (algebra.algebra_map_submonoid S M) Sₘ] variables [algebra Rₘ Sₘ] [algebra R Sₘ] [is_scalar_tower R Rₘ Sₘ] [is_scalar_tower R S Sₘ] +include M /-- Let `S` be an extension of `R` and `Rₘ Sₘ` be localizations at `M` of `R S` respectively. Then the norm of `a : Sₘ` over `Rₘ` is the norm of `a : S` over `R` if `S` is free as `R`-module. -/ -lemma algebra.norm_localization [module.free R S] [module.finite R S] - (a : S) (hM : algebra.algebra_map_submonoid S M ≤ S⁰) : +lemma algebra.norm_localization [module.free R S] [module.finite R S] (a : S) : algebra.norm Rₘ (algebra_map S Sₘ a) = algebra_map R Rₘ (algebra.norm R a) := begin casesI subsingleton_or_nontrivial R, @@ -45,7 +45,7 @@ begin simp }, let b := module.free.choose_basis R S, letI := classical.dec_eq (module.free.choose_basis_index R S), - rw [algebra.norm_eq_matrix_det (b.localization_localization Rₘ M Sₘ hM), + rw [algebra.norm_eq_matrix_det (b.localization_localization Rₘ M Sₘ), algebra.norm_eq_matrix_det b, ring_hom.map_det], congr, ext i j, From da1d134ab55eb58347924920695d8200f4740694 Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Thu, 2 Feb 2023 11:40:06 +0000 Subject: [PATCH 0413/1291] feat(analysis/special_functions/pow): strengthen cpow integrability lemmas (#18354) This adds some slightly stronger results about continuity / integrability of `x -> x^s` for complex `s`. --- src/analysis/special_functions/integrals.lean | 184 +++++++++++++----- src/analysis/special_functions/pow.lean | 88 ++++++--- src/analysis/special_functions/pow_deriv.lean | 38 ++++ 3 files changed, 236 insertions(+), 74 deletions(-) diff --git a/src/analysis/special_functions/integrals.lean b/src/analysis/special_functions/integrals.lean index 6b6b24c57e75c..c657a727b5564 100644 --- a/src/analysis/special_functions/integrals.lean +++ b/src/analysis/special_functions/integrals.lean @@ -48,11 +48,14 @@ lemma interval_integrable_zpow {n : ℤ} (h : 0 ≤ n ∨ (0 : ℝ) ∉ [a, b]) interval_integrable (λ x, x ^ n) μ a b := (continuous_on_id.zpow₀ n $ λ x hx, h.symm.imp (ne_of_mem_of_not_mem hx) id).interval_integrable +/-- See `interval_integrable_rpow'` for a version with a weaker hypothesis on `r`, but assuming the +measure is volume. -/ lemma interval_integrable_rpow {r : ℝ} (h : 0 ≤ r ∨ (0 : ℝ) ∉ [a, b]) : interval_integrable (λ x, x ^ r) μ a b := (continuous_on_id.rpow_const $ λ x hx, h.symm.imp (ne_of_mem_of_not_mem hx) id).interval_integrable -/-- Alternative version with a weaker hypothesis on `r`, but assuming the measure is volume. -/ +/-- See `interval_integrable_rpow` for a version applying to any locally finite measure, but with a +stronger hypothesis on `r`. -/ lemma interval_integrable_rpow' {r : ℝ} (h : -1 < r) : interval_integrable (λ x, x ^ r) volume a b := begin @@ -78,13 +81,92 @@ begin rpow_def_of_pos hx.1, rpow_def_of_neg (by linarith [hx.1] : -x < 0)], } end -lemma interval_integrable_cpow {r : ℂ} (ha : 0 < a) (hb : 0 < b) : - interval_integrable (λ x : ℝ, (x : ℂ) ^ r) volume a b := +/-- See `interval_integrable_cpow'` for a version with a weaker hypothesis on `r`, but assuming the +measure is volume. -/ +lemma interval_integrable_cpow {r : ℂ} (h : 0 ≤ r.re ∨ (0 : ℝ) ∉ [a, b]) : + interval_integrable (λ x : ℝ, (x : ℂ) ^ r) μ a b := begin - refine (complex.continuous_of_real.continuous_on.cpow_const _).interval_integrable, - intros c hc, - left, - exact_mod_cast lt_of_lt_of_le (lt_min ha hb) hc.left, + by_cases h2 : (0:ℝ) ∉ [a,b], + { -- Easy case #1: 0 ∉ [a, b] -- use continuity. + refine (continuous_at.continuous_on (λ x hx, _)).interval_integrable, + exact complex.continuous_at_of_real_cpow_const _ _ (or.inr $ ne_of_mem_of_not_mem hx h2) }, + rw [eq_false_intro h2, or_false] at h, + rcases lt_or_eq_of_le h with h'|h', + { -- Easy case #2: 0 < re r -- again use continuity + exact (complex.continuous_of_real_cpow_const h').interval_integrable _ _ }, + -- Now the hard case: re r = 0 and 0 is in the interval. + refine (interval_integrable.interval_integrable_norm_iff _).mp _, + { refine (measurable_of_continuous_on_compl_singleton (0:ℝ) _).ae_strongly_measurable, + exact continuous_at.continuous_on + (λ x hx, complex.continuous_at_of_real_cpow_const x r (or.inr hx)) }, + -- reduce to case of integral over `[0, c]` + suffices : ∀ (c : ℝ), interval_integrable (λ x : ℝ, ‖↑x ^ r‖) μ 0 c, + from (this a).symm.trans (this b), + intro c, + rcases le_or_lt 0 c with hc | hc, + { -- case `0 ≤ c`: integrand is identically 1 + have : interval_integrable (λ x, 1 : ℝ → ℝ) μ 0 c, + from interval_integrable_const, + rw interval_integrable_iff_integrable_Ioc_of_le hc at this ⊢, + refine integrable_on.congr_fun this (λ x hx, _) measurable_set_Ioc, + dsimp only, + rw [complex.norm_eq_abs, complex.abs_cpow_eq_rpow_re_of_pos hx.1, ←h', rpow_zero], }, + { -- case `c < 0`: integrand is identically constant, *except* at `x = 0` if `r ≠ 0`. + apply interval_integrable.symm, + rw [interval_integrable_iff_integrable_Ioc_of_le hc.le], + have : Ioc c 0 = Ioo c 0 ∪ {(0:ℝ)}, + { rw [←Ioo_union_Icc_eq_Ioc hc (le_refl 0), ←Icc_def], + simp_rw [←le_antisymm_iff, set_of_eq_eq_singleton'] }, + rw [this, integrable_on_union, and.comm], split, + { refine integrable_on_singleton_iff.mpr (or.inr _), + exact is_finite_measure_on_compacts_of_is_locally_finite_measure.lt_top_of_is_compact + is_compact_singleton }, + { have : ∀ (x : ℝ), x ∈ Ioo c 0 → ‖complex.exp (↑π * complex.I * r)‖ = ‖(x:ℂ) ^ r‖, + { intros x hx, + rw [complex.of_real_cpow_of_nonpos hx.2.le, norm_mul, ←complex.of_real_neg, + complex.norm_eq_abs (_ ^ _), complex.abs_cpow_eq_rpow_re_of_pos (neg_pos.mpr hx.2), + ←h', rpow_zero, one_mul] }, + refine integrable_on.congr_fun _ this measurable_set_Ioo, + rw integrable_on_const, + refine or.inr ((measure_mono set.Ioo_subset_Icc_self).trans_lt _), + exact is_finite_measure_on_compacts_of_is_locally_finite_measure.lt_top_of_is_compact + is_compact_Icc } }, +end + +/-- See `interval_integrable_cpow` for a version applying to any locally finite measure, but with a +stronger hypothesis on `r`. -/ +lemma interval_integrable_cpow' {r : ℂ} (h : -1 < r.re) : + interval_integrable (λ x:ℝ, (x:ℂ) ^ r) volume a b := +begin + suffices : ∀ (c : ℝ), interval_integrable (λ x, (x : ℂ) ^ r) volume 0 c, + { exact interval_integrable.trans (this a).symm (this b) }, + have : ∀ (c : ℝ), (0 ≤ c) → interval_integrable (λ x, (x : ℂ) ^ r) volume 0 c, + { intros c hc, + rw ←interval_integrable.interval_integrable_norm_iff, + { rw interval_integrable_iff, + apply integrable_on.congr_fun, + { rw ←interval_integrable_iff, exact (interval_integral.interval_integrable_rpow' h) }, + { intros x hx, + rw uIoc_of_le hc at hx, + dsimp only, + rw [complex.norm_eq_abs, complex.abs_cpow_eq_rpow_re_of_pos hx.1] }, + { exact measurable_set_uIoc } }, + { refine continuous_on.ae_strongly_measurable _ measurable_set_uIoc, + refine continuous_at.continuous_on (λ x hx, _), + rw uIoc_of_le hc at hx, + refine (continuous_at_cpow_const (or.inl _)).comp complex.continuous_of_real.continuous_at, + rw complex.of_real_re, + exact hx.1 } }, + intro c, rcases le_total 0 c with hc | hc, + { exact this c hc }, + { rw [interval_integrable.iff_comp_neg, neg_zero], + have m := (this (-c) (by linarith)).const_mul (complex.exp (π * complex.I * r)), + rw [interval_integrable_iff, uIoc_of_le (by linarith : 0 ≤ -c)] at m ⊢, + refine m.congr_fun (λ x hx, _) measurable_set_Ioc, + dsimp only, + have : -x ≤ 0, by linarith [hx.1], + rw [complex.of_real_cpow_of_nonpos this, mul_comm], + simp } end @[simp] @@ -192,56 +274,58 @@ open interval_integral /-! ### Integrals of simple functions -/ -lemma integral_rpow {r : ℝ} (h : -1 < r ∨ (r ≠ -1 ∧ (0 : ℝ) ∉ [a, b])) : - ∫ x in a..b, x ^ r = (b ^ (r + 1) - a ^ (r + 1)) / (r + 1) := +lemma integral_cpow {r : ℂ} (h : -1 < r.re ∨ (r ≠ -1 ∧ (0 : ℝ) ∉ [a, b])) : + ∫ (x : ℝ) in a..b, (x : ℂ) ^ r = (b ^ (r + 1) - a ^ (r + 1)) / (r + 1) := begin rw sub_div, - have hderiv : ∀ (x : ℝ), x ≠ 0 → has_deriv_at (λ x : ℝ, x ^ (r + 1) / (r + 1)) (x ^ r) x, - { intros x hx, - convert (real.has_deriv_at_rpow_const (or.inl hx)).div_const (r + 1), - rw [add_sub_cancel, mul_div_cancel_left], - contrapose! h, rw ←eq_neg_iff_add_eq_zero at h, rw h, tauto, }, - cases h, - { suffices : ∀ (c : ℝ), ∫ x in 0..c, x ^ r = c ^ (r + 1) / (r + 1), - { rw [←integral_add_adjacent_intervals - (interval_integrable_rpow' h) (interval_integrable_rpow' h), this b], - have t := this a, rw integral_symm at t, apply_fun (λ x, -x) at t, rw neg_neg at t, - rw t, ring }, - intro c, rcases le_total 0 c with hc|hc, - { convert integral_eq_sub_of_has_deriv_at_of_le hc _ (λ x hx, hderiv x hx.1.ne') _, - { rw zero_rpow, ring, linarith,}, - { apply continuous_at.continuous_on, intros x hx, - refine (continuous_at_id.rpow_const _).div_const, right, linarith,}, - apply interval_integrable_rpow' h }, - { rw integral_symm, symmetry, rw eq_neg_iff_eq_neg, - convert integral_eq_sub_of_has_deriv_at_of_le hc _ (λ x hx, hderiv x hx.2.ne) _, - { rw zero_rpow, ring, linarith }, - { apply continuous_at.continuous_on, intros x hx, - refine (continuous_at_id.rpow_const _).div_const, right, linarith }, - apply interval_integrable_rpow' h, } }, - { have hderiv' : ∀ (x : ℝ), x ∈ [a, b] → has_deriv_at (λ x : ℝ, x ^ (r + 1) / (r + 1)) (x ^ r) x, - { intros x hx, apply hderiv x, exact ne_of_mem_of_not_mem hx h.2 }, - exact integral_eq_sub_of_has_deriv_at hderiv' (interval_integrable_rpow (or.inr h.2)) }, + have hr : r + 1 ≠ 0, + { cases h, + { apply_fun complex.re, + rw [complex.add_re, complex.one_re, complex.zero_re, ne.def, add_eq_zero_iff_eq_neg], + exact h.ne' }, + { rw [ne.def, ←add_eq_zero_iff_eq_neg] at h, exact h.1 } }, + by_cases hab : (0:ℝ) ∉ [a, b], + { refine integral_eq_sub_of_has_deriv_at (λ x hx, _) (interval_integrable_cpow $ or.inr hab), + refine has_deriv_at_of_real_cpow (ne_of_mem_of_not_mem hx hab) _, + contrapose! hr, rwa add_eq_zero_iff_eq_neg }, + replace h : -1 < r.re, by tauto, + suffices : ∀ (c : ℝ), ∫ (x : ℝ) in 0..c, (x : ℂ) ^ r = + c ^ (r + 1) / (r + 1) - 0 ^ (r + 1) / (r + 1), + { rw [←integral_add_adjacent_intervals (@interval_integrable_cpow' a 0 r h) + (@interval_integrable_cpow' 0 b r h), integral_symm, this a, this b, complex.zero_cpow hr], + ring }, + intro c, + apply integral_eq_sub_of_has_deriv_right, + { refine (complex.continuous_of_real_cpow_const _).div_const.continuous_on, + rwa [complex.add_re, complex.one_re, ←neg_lt_iff_pos_add] }, + { refine λ x hx, (has_deriv_at_of_real_cpow _ _).has_deriv_within_at, + { rcases le_total c 0 with hc | hc, + { rw max_eq_left hc at hx, exact hx.2.ne }, { rw min_eq_left hc at hx, exact hx.1.ne' } }, + { contrapose! hr, rw hr, ring } }, + { exact interval_integrable_cpow' h } end -lemma integral_cpow {r : ℂ} (ha : 0 < a) (hb : 0 < b) (hr : r ≠ -1) : - ∫ (x : ℝ) in a..b, (x : ℂ) ^ r = (b ^ (r + 1) - a ^ (r + 1)) / (r + 1) := +lemma integral_rpow {r : ℝ} (h : -1 < r ∨ (r ≠ -1 ∧ (0 : ℝ) ∉ [a, b])) : + ∫ x in a..b, x ^ r = (b ^ (r + 1) - a ^ (r + 1)) / (r + 1) := begin - rw sub_div, - suffices : ∀ x ∈ set.uIcc a b, has_deriv_at (λ z : ℂ, z ^ (r + 1) / (r + 1)) (x ^ r) x, - { exact integral_eq_sub_of_has_deriv_at - (λ x hx, (this x hx).comp_of_real) (interval_integrable_cpow ha hb) }, - intros x hx, - have hx' : 0 < (x : ℂ).re ∨ (x : ℂ).im ≠ 0, - { left, - norm_cast, - calc 0 < min a b : lt_min ha hb ... ≤ x : hx.left, }, - convert ((has_deriv_at_id (x:ℂ)).cpow_const hx').div_const (r + 1), - simp only [id.def, add_sub_cancel, mul_one], rw [mul_comm, mul_div_cancel], - contrapose! hr, rwa add_eq_zero_iff_eq_neg at hr, + have h' : -1 < (r:ℂ).re ∨ (r:ℂ) ≠ -1 ∧ (0:ℝ) ∉ [a, b], + { cases h, + { left, rwa complex.of_real_re }, + { right, rwa [←complex.of_real_one, ←complex.of_real_neg, ne.def, complex.of_real_inj] } }, + have : ∫ x in a..b, (x:ℂ) ^ (r :ℂ) = ((b:ℂ) ^ (r + 1 : ℂ) - (a:ℂ) ^ (r + 1 : ℂ)) / (r + 1), + from integral_cpow h', + apply_fun complex.re at this, convert this, + { simp_rw [interval_integral_eq_integral_uIoc, complex.real_smul, complex.of_real_mul_re], + { change complex.re with is_R_or_C.re, + rw ←integral_re, refl, + refine interval_integrable_iff.mp _, + cases h', + { exact interval_integrable_cpow' h' }, { exact interval_integrable_cpow (or.inr h'.2) } } }, + { rw (by push_cast : ((r:ℂ) + 1) = ((r + 1 : ℝ) : ℂ)), + simp_rw [div_eq_inv_mul, ←complex.of_real_inv, complex.of_real_mul_re, complex.sub_re], + refl } end - lemma integral_zpow {n : ℤ} (h : 0 ≤ n ∨ n ≠ -1 ∧ (0 : ℝ) ∉ [a, b]) : ∫ x in a..b, x ^ n = (b ^ (n + 1) - a ^ (n + 1)) / (n + 1) := begin diff --git a/src/analysis/special_functions/pow.lean b/src/analysis/special_functions/pow.lean index 5570344bf000c..c7cb1b7f2eaf0 100644 --- a/src/analysis/special_functions/pow.lean +++ b/src/analysis/special_functions/pow.lean @@ -134,6 +134,21 @@ begin ... ≤ π * n : mul_le_mul_of_nonneg_left hn1 real.pi_pos.le } end +lemma mul_cpow_of_real_nonneg {a b : ℝ} (ha : 0 ≤ a) (hb : 0 ≤ b) (r : ℂ) : + ((a : ℂ) * (b : ℂ)) ^ r = (a : ℂ) ^ r * (b : ℂ) ^ r := +begin + rcases eq_or_ne r 0 with rfl | hr, + { simp only [cpow_zero, mul_one] }, + rcases eq_or_lt_of_le ha with rfl | ha', + { rw [of_real_zero, zero_mul, zero_cpow hr, zero_mul] }, + rcases eq_or_lt_of_le hb with rfl | hb', + { rw [of_real_zero, mul_zero, zero_cpow hr, mul_zero] }, + have ha'' : (a : ℂ) ≠ 0 := of_real_ne_zero.mpr ha'.ne', + have hb'' : (b : ℂ) ≠ 0 := of_real_ne_zero.mpr hb'.ne', + rw [cpow_def_of_ne_zero (mul_ne_zero ha'' hb''), log_of_real_mul ha' hb'', of_real_log ha, + add_mul, exp_add, ←cpow_def_of_ne_zero ha'', ←cpow_def_of_ne_zero hb''] +end + end complex section lim @@ -143,18 +158,18 @@ open complex variables {α : Type*} lemma zero_cpow_eq_nhds {b : ℂ} (hb : b ≠ 0) : - (0 : ℂ).cpow =ᶠ[𝓝 b] 0 := + (λ (x : ℂ), (0 : ℂ) ^ x) =ᶠ[𝓝 b] 0 := begin suffices : ∀ᶠ (x : ℂ) in (𝓝 b), x ≠ 0, - from this.mono (λ x hx, by rw [cpow_eq_pow, zero_cpow hx, pi.zero_apply]), + from this.mono (λ x hx, by { dsimp only, rw [zero_cpow hx, pi.zero_apply]} ), exact is_open.eventually_mem is_open_ne hb, end lemma cpow_eq_nhds {a b : ℂ} (ha : a ≠ 0) : - (λ x, x.cpow b) =ᶠ[𝓝 a] λ x, exp (log x * b) := + (λ x, x ^ b) =ᶠ[𝓝 a] λ x, exp (log x * b) := begin suffices : ∀ᶠ (x : ℂ) in (𝓝 a), x ≠ 0, - from this.mono (λ x hx, by { dsimp only, rw [cpow_eq_pow, cpow_def_of_ne_zero hx], }), + from this.mono (λ x hx, by { dsimp only, rw [cpow_def_of_ne_zero hx], }), exact is_open.eventually_mem is_open_ne ha, end @@ -169,15 +184,17 @@ begin exact is_closed_eq continuous_fst continuous_const, end -lemma continuous_at_const_cpow {a b : ℂ} (ha : a ≠ 0) : continuous_at (cpow a) b := +/- Continuity of `λ x, a ^ x`: union of these two lemmas is optimal. -/ + +lemma continuous_at_const_cpow {a b : ℂ} (ha : a ≠ 0) : continuous_at (λ x, a ^ x) b := begin - have cpow_eq : cpow a = λ b, exp (log a * b), - by { ext1 b, rw [cpow_eq_pow, cpow_def_of_ne_zero ha], }, + have cpow_eq : (λ x:ℂ, a ^ x) = λ x, exp (log a * x), + by { ext1 b, rw [cpow_def_of_ne_zero ha], }, rw cpow_eq, exact continuous_exp.continuous_at.comp (continuous_at.mul continuous_at_const continuous_at_id), end -lemma continuous_at_const_cpow' {a b : ℂ} (h : b ≠ 0) : continuous_at (cpow a) b := +lemma continuous_at_const_cpow' {a b : ℂ} (h : b ≠ 0) : continuous_at (λ x, a ^ x) b := begin by_cases ha : a = 0, { rw [ha, continuous_at_congr (zero_cpow_eq_nhds h)], exact continuous_at_const, }, @@ -1211,7 +1228,7 @@ end limits namespace complex -/-- See also `complex.continuous_at_cpow` and `complex.continuous_at_cpow_of_re_pos`. -/ +/-- See also `continuous_at_cpow` and `complex.continuous_at_cpow_of_re_pos`. -/ lemma continuous_at_cpow_zero_of_re_pos {z : ℂ} (hz : 0 < z.re) : continuous_at (λ x : ℂ × ℂ, x.1 ^ x.2) (0, z) := begin @@ -1232,7 +1249,7 @@ begin (_root_.abs_nonneg _) real.pi_pos.le } end -/-- See also `complex.continuous_at_cpow` for a version that assumes `p.1 ≠ 0` but makes no +/-- See also `continuous_at_cpow` for a version that assumes `p.1 ≠ 0` but makes no assumptions about `p.2`. -/ lemma continuous_at_cpow_of_re_pos {p : ℂ × ℂ} (h₁ : 0 ≤ p.1.re ∨ p.1.im ≠ 0) (h₂ : 0 < p.2.re) : continuous_at (λ x : ℂ × ℂ, x.1 ^ x.2) p := @@ -1243,27 +1260,50 @@ begin exacts [continuous_at_cpow h₁, continuous_at_cpow_zero_of_re_pos h₂] end -/-- See also `complex.continuous_at_cpow_const` for a version that assumes `z ≠ 0` but makes no +/-- See also `continuous_at_cpow_const` for a version that assumes `z ≠ 0` but makes no assumptions about `w`. -/ lemma continuous_at_cpow_const_of_re_pos {z w : ℂ} (hz : 0 ≤ re z ∨ im z ≠ 0) (hw : 0 < re w) : continuous_at (λ x, x ^ w) z := tendsto.comp (@continuous_at_cpow_of_re_pos (z, w) hz hw) (continuous_at_id.prod continuous_at_const) +/-- Continuity of `(x, y) ↦ x ^ y` as a function on `ℝ × ℂ`. -/ +lemma continuous_at_of_real_cpow (x : ℝ) (y : ℂ) (h : 0 < y.re ∨ x ≠ 0) : + continuous_at (λ p, ↑p.1 ^ p.2 : ℝ × ℂ → ℂ) (x, y) := +begin + rcases lt_trichotomy 0 x with hx | rfl | hx, + { -- x > 0 : easy case + have : continuous_at (λ p, ⟨↑p.1, p.2⟩ : ℝ × ℂ → ℂ × ℂ) (x, y), + from continuous_of_real.continuous_at.prod_map continuous_at_id, + refine (continuous_at_cpow (or.inl _)).comp this, + rwa of_real_re }, + { -- x = 0 : reduce to continuous_at_cpow_zero_of_re_pos + have A : continuous_at (λ p, p.1 ^ p.2 : ℂ × ℂ → ℂ) ⟨↑(0:ℝ), y⟩, + { rw of_real_zero, + apply continuous_at_cpow_zero_of_re_pos, + tauto }, + have B : continuous_at (λ p, ⟨↑p.1, p.2⟩ : ℝ × ℂ → ℂ × ℂ) ⟨0, y⟩, + from continuous_of_real.continuous_at.prod_map continuous_at_id, + exact @continuous_at.comp (ℝ × ℂ) (ℂ × ℂ) ℂ _ _ _ _ (λ p, ⟨↑p.1, p.2⟩) ⟨0, y⟩ A B }, + { -- x < 0 : difficult case + suffices : continuous_at (λ p, (-↑p.1) ^ p.2 * exp (π * I * p.2) : ℝ × ℂ → ℂ) (x, y), + { refine this.congr (eventually_of_mem (prod_mem_nhds (Iio_mem_nhds hx) univ_mem) _), + exact λ p hp, (of_real_cpow_of_nonpos (le_of_lt hp.1) p.2).symm }, + have A : continuous_at (λ p, ⟨-↑p.1, p.2⟩ : ℝ × ℂ → ℂ × ℂ) (x, y), + from continuous_at.prod_map (continuous_of_real.continuous_at.neg) continuous_at_id, + apply continuous_at.mul, + { refine (continuous_at_cpow (or.inl _)).comp A, + rwa [neg_re, of_real_re, neg_pos] }, + { exact (continuous_exp.comp (continuous_const.mul continuous_snd)).continuous_at } }, +end + +lemma continuous_at_of_real_cpow_const (x : ℝ) (y : ℂ) (h : 0 < y.re ∨ x ≠ 0) : + continuous_at (λ a, a ^ y : ℝ → ℂ) x := +@continuous_at.comp _ _ _ _ _ _ _ _ x (continuous_at_of_real_cpow x y h) + (continuous_id.prod_mk continuous_const).continuous_at + lemma continuous_of_real_cpow_const {y : ℂ} (hs : 0 < y.re) : continuous (λ x, x ^ y : ℝ → ℂ) := -begin - rw continuous_iff_continuous_at, intro x, - cases le_or_lt 0 x with hx hx, - { refine (continuous_at_cpow_const_of_re_pos _ hs).comp continuous_of_real.continuous_at, - exact or.inl hx }, - { suffices : continuous_on (λ x, x ^ y : ℝ → ℂ) (set.Iio 0), - from continuous_on.continuous_at this (Iio_mem_nhds hx), - have : eq_on (λ x, x ^ y : ℝ → ℂ) (λ x, ((-x) : ℂ) ^ y * exp (π * I * y)) (set.Iio 0), - from λ y hy, of_real_cpow_of_nonpos (le_of_lt hy) _, - refine (continuous_on.mul (λ y hy, _) continuous_on_const).congr this, - refine continuous_of_real.continuous_within_at.neg.cpow continuous_within_at_const _, - left, simpa using hy } -end +continuous_iff_continuous_at.mpr (λ x, continuous_at_of_real_cpow_const x y (or.inl hs)) end complex diff --git a/src/analysis/special_functions/pow_deriv.lean b/src/analysis/special_functions/pow_deriv.lean index 6793eae1134f3..c785353eef25c 100644 --- a/src/analysis/special_functions/pow_deriv.lean +++ b/src/analysis/special_functions/pow_deriv.lean @@ -187,6 +187,44 @@ lemma has_deriv_within_at.cpow_const (hf : has_deriv_within_at f f' s x) has_deriv_within_at (λ x, f x ^ c) (c * f x ^ (c - 1) * f') s x := (complex.has_strict_deriv_at_cpow_const h0).has_deriv_at.comp_has_deriv_within_at x hf +/-- Although `λ x, x ^ r` for fixed `r` is *not* complex-differentiable along the negative real +line, it is still real-differentiable, and the derivative is what one would formally expect. -/ +lemma has_deriv_at_of_real_cpow {x : ℝ} (hx : x ≠ 0) {r : ℂ} (hr : r ≠ -1) : + has_deriv_at (λ y:ℝ, (y:ℂ) ^ (r + 1) / (r + 1)) (x ^ r) x := +begin + rw [ne.def, ←add_eq_zero_iff_eq_neg, ←ne.def] at hr, + rcases lt_or_gt_of_ne hx.symm with hx | hx, + { -- easy case : `0 < x` + convert (((has_deriv_at_id (x:ℂ)).cpow_const _).div_const (r + 1)).comp_of_real, + { rw [add_sub_cancel, id.def, mul_one, mul_comm, mul_div_cancel _ hr] }, + { rw [id.def, of_real_re], exact or.inl hx } }, + { -- harder case : `x < 0` + have : ∀ᶠ (y:ℝ) in nhds x, (y:ℂ) ^ (r + 1) / (r + 1) = + (-y:ℂ) ^ (r + 1) * exp (π * I * (r + 1)) / (r + 1), + { refine filter.eventually_of_mem (Iio_mem_nhds hx) (λ y hy, _), + rw of_real_cpow_of_nonpos (le_of_lt hy) }, + refine has_deriv_at.congr_of_eventually_eq _ this, + rw of_real_cpow_of_nonpos (le_of_lt hx), + suffices : has_deriv_at (λ (y : ℝ), (-↑y) ^ (r + 1) * exp (↑π * I * (r + 1))) + ((r + 1) * (-↑x) ^ r * exp (↑π * I * r)) x, + { convert this.div_const (r + 1) using 1, + conv_rhs { rw [mul_assoc, mul_comm, mul_div_cancel _ hr] } }, + rw [mul_add ((π:ℂ) * _), mul_one, exp_add, exp_pi_mul_I, + mul_comm (_ : ℂ) (-1 : ℂ), neg_one_mul], + simp_rw [mul_neg, ←neg_mul, ←of_real_neg], + suffices : has_deriv_at (λ (y : ℝ), (↑-y) ^ (r + 1)) (-(r + 1) * (↑-x) ^ r) x, + { convert this.neg.mul_const _, ring }, + suffices : has_deriv_at (λ (y : ℝ), (↑y) ^ (r + 1)) ((r + 1) * (↑-x) ^ r) (-x), + { convert @has_deriv_at.scomp ℝ _ ℂ _ _ x ℝ _ _ _ _ _ _ _ _ this (has_deriv_at_neg x) using 1, + rw [real_smul, of_real_neg 1, of_real_one], ring }, + suffices : has_deriv_at (λ (y : ℂ), y ^ (r + 1)) ((r + 1) * (↑-x) ^ r) (↑-x), + { exact this.comp_of_real }, + conv in ((↑_) ^ _) { rw (by ring : r = (r + 1) - 1) }, + convert (has_deriv_at_id ((-x : ℝ) : ℂ)).cpow_const _ using 1, + { simp }, + { left, rwa [id.def, of_real_re, neg_pos] } }, +end + end deriv namespace real From 8a035e9a1adeb17db35b03a253e8ef0ecdb6e4d8 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 2 Feb 2023 14:06:14 +0000 Subject: [PATCH 0414/1291] feat(counterexamples/quadratic_form): symmetric bilinear forms in char 2 do not always inject into quadratic forms (#18146) --- counterexamples/quadratic_form.lean | 53 +++++++++++++++++++++++++++++ 1 file changed, 53 insertions(+) create mode 100644 counterexamples/quadratic_form.lean diff --git a/counterexamples/quadratic_form.lean b/counterexamples/quadratic_form.lean new file mode 100644 index 0000000000000..20af196c19f52 --- /dev/null +++ b/counterexamples/quadratic_form.lean @@ -0,0 +1,53 @@ +/- +Copyright (c) 2023 Eric Wieser. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Eric Wieser +-/ +import linear_algebra.quadratic_form.basic +import algebra.char_p.two +import data.zmod.basic + +/-! +# `quadratic_form R M` and `subtype bilin_form.is_symm` are distinct notions in characteristic 2 + +The main result of this file is `bilin_form.not_inj_on_to_quadratic_form_is_symm`. + +The counterexample we use is $B (x, y) (x', y') ↦ xy' + x'y$ where `x y x' y' : zmod 2`. +-/ + +variables (F : Type*) [nontrivial F] [comm_ring F] [char_p F 2] + +open bilin_form + +/-- The bilinear form we will use as a counterexample, over some field `F` of characteristic two. -/ +def B : bilin_form F (F × F) := +bilin_form.lin_mul_lin (linear_map.fst _ _ _) (linear_map.snd _ _ _) + + bilin_form.lin_mul_lin (linear_map.snd _ _ _) (linear_map.fst _ _ _) + +@[simp] +lemma B_apply (x y : F × F) : B F x y = x.1 * y.2 + x.2 * y.1 := rfl + +lemma is_symm_B : (B F).is_symm := λ x y, by simp [mul_comm, add_comm] + +lemma is_alt_B : (B F).is_alt := λ x, by simp [mul_comm, char_two.add_self_eq_zero (x.1 * x.2)] + +lemma B_ne_zero : B F ≠ 0 := λ h, by simpa using bilin_form.congr_fun h (1, 0) (1, 1) + +/-- `bilin_form.to_quadratic_form` is not injective on symmetric bilinear forms. + +This disproves a weaker version of `quadratic_form.associated_left_inverse`. +-/ +lemma {u} bilin_form.not_inj_on_to_quadratic_form_is_symm : + ¬∀ {R M : Type u} [semiring R] [add_comm_monoid M], + by exactI ∀ [module R M], + by exactI set.inj_on + (to_quadratic_form : bilin_form R M → quadratic_form R M) + { B | B.is_symm }:= +begin + intro h, + let F := ulift.{u} (zmod 2), + apply B_ne_zero F, + apply h (is_symm_B F) (is_symm_zero), + rw [bilin_form.to_quadratic_form_zero, bilin_form.to_quadratic_form_eq_zero], + exact is_alt_B F, +end From 2705404e701abc6b3127da906f40bae062a169c9 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 2 Feb 2023 17:00:14 +0000 Subject: [PATCH 0415/1291] refactor(linear_algebra/basic): extract content about linear_map.general_linear_group (#18345) --- .../affine_space/affine_equiv.lean | 1 + src/linear_algebra/basic.lean | 53 -------------- src/linear_algebra/determinant.lean | 1 + src/linear_algebra/eigenspace.lean | 1 + src/linear_algebra/general_linear_group.lean | 73 +++++++++++++++++++ .../matrix/general_linear_group.lean | 1 + .../matrix/special_linear_group.lean | 1 + src/linear_algebra/unitary_group.lean | 1 + src/number_theory/modular.lean | 1 + 9 files changed, 80 insertions(+), 53 deletions(-) create mode 100644 src/linear_algebra/general_linear_group.lean diff --git a/src/linear_algebra/affine_space/affine_equiv.lean b/src/linear_algebra/affine_space/affine_equiv.lean index 72b42e85a3e29..42559784f88d9 100644 --- a/src/linear_algebra/affine_space/affine_equiv.lean +++ b/src/linear_algebra/affine_space/affine_equiv.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury G. Kudryashov -/ import linear_algebra.affine_space.affine_map +import linear_algebra.general_linear_group import algebra.invertible /-! diff --git a/src/linear_algebra/basic.lean b/src/linear_algebra/basic.lean index 00493cd4349f5..dfc8322aa2fbc 100644 --- a/src/linear_algebra/basic.lean +++ b/src/linear_algebra/basic.lean @@ -25,7 +25,6 @@ Many of the relevant definitions, including `module`, `submodule`, and `linear_m * Many constructors for (semi)linear maps * The kernel `ker` and range `range` of a linear map are submodules of the domain and codomain respectively. -* The general linear group is defined to be the group of invertible linear maps from `M` to itself. See `linear_algebra.span` for the span of a set (as a submodule), and `linear_algebra.quotient` for quotients by submodules. @@ -2161,55 +2160,3 @@ rfl end linear_equiv end fun_left - -namespace linear_map - -variables [semiring R] [add_comm_monoid M] [module R M] -variables (R M) - -/-- The group of invertible linear maps from `M` to itself -/ -@[reducible] def general_linear_group := (M →ₗ[R] M)ˣ - -namespace general_linear_group -variables {R M} - -instance : has_coe_to_fun (general_linear_group R M) (λ _, M → M) := by apply_instance - -/-- An invertible linear map `f` determines an equivalence from `M` to itself. -/ -def to_linear_equiv (f : general_linear_group R M) : (M ≃ₗ[R] M) := -{ inv_fun := f.inv.to_fun, - left_inv := λ m, show (f.inv * f.val) m = m, - by erw f.inv_val; simp, - right_inv := λ m, show (f.val * f.inv) m = m, - by erw f.val_inv; simp, - ..f.val } - -/-- An equivalence from `M` to itself determines an invertible linear map. -/ -def of_linear_equiv (f : (M ≃ₗ[R] M)) : general_linear_group R M := -{ val := f, - inv := (f.symm : M →ₗ[R] M), - val_inv := linear_map.ext $ λ _, f.apply_symm_apply _, - inv_val := linear_map.ext $ λ _, f.symm_apply_apply _ } - -variables (R M) - -/-- The general linear group on `R` and `M` is multiplicatively equivalent to the type of linear -equivalences between `M` and itself. -/ -def general_linear_equiv : general_linear_group R M ≃* (M ≃ₗ[R] M) := -{ to_fun := to_linear_equiv, - inv_fun := of_linear_equiv, - left_inv := λ f, by { ext, refl }, - right_inv := λ f, by { ext, refl }, - map_mul' := λ x y, by {ext, refl} } - -@[simp] lemma general_linear_equiv_to_linear_map (f : general_linear_group R M) : - (general_linear_equiv R M f : M →ₗ[R] M) = f := -by {ext, refl} - -@[simp] lemma coe_fn_general_linear_equiv (f : general_linear_group R M) : - ⇑(general_linear_equiv R M f) = (f : M → M) := -rfl - -end general_linear_group - -end linear_map diff --git a/src/linear_algebra/determinant.lean b/src/linear_algebra/determinant.lean index a10bd7d34f0f1..1623a1ec203b5 100644 --- a/src/linear_algebra/determinant.lean +++ b/src/linear_algebra/determinant.lean @@ -3,6 +3,7 @@ Copyright (c) 2019 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Patrick Massot, Casper Putz, Anne Baanen -/ +import linear_algebra.general_linear_group import linear_algebra.matrix.reindex import tactic.field_simp import linear_algebra.matrix.nonsingular_inverse diff --git a/src/linear_algebra/eigenspace.lean b/src/linear_algebra/eigenspace.lean index ecdc77cacc826..59b617cdef3d6 100644 --- a/src/linear_algebra/eigenspace.lean +++ b/src/linear_algebra/eigenspace.lean @@ -7,6 +7,7 @@ Authors: Alexander Bentkamp import algebra.algebra.spectrum import order.hom.basic import linear_algebra.free_module.finite.basic +import linear_algebra.general_linear_group /-! # Eigenvectors and eigenvalues diff --git a/src/linear_algebra/general_linear_group.lean b/src/linear_algebra/general_linear_group.lean new file mode 100644 index 0000000000000..d392c759b6dd0 --- /dev/null +++ b/src/linear_algebra/general_linear_group.lean @@ -0,0 +1,73 @@ +/- +Copyright (c) 2019 Johan Commelin. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johan Commelin +-/ +import algebra.module.equiv + +/-! +# The general linear group of linear maps + +The general linear group is defined to be the group of invertible linear maps from `M` to itself. + +See also `matrix.general_linear_group` + +## Main definitions + +* `linear_map.general_linear_group` + +-/ + +variables (R M : Type*) + +namespace linear_map + +variables [semiring R] [add_comm_monoid M] [module R M] +variables (R M) + +/-- The group of invertible linear maps from `M` to itself -/ +@[reducible] def general_linear_group := (M →ₗ[R] M)ˣ + +namespace general_linear_group +variables {R M} + +instance : has_coe_to_fun (general_linear_group R M) (λ _, M → M) := by apply_instance + +/-- An invertible linear map `f` determines an equivalence from `M` to itself. -/ +def to_linear_equiv (f : general_linear_group R M) : (M ≃ₗ[R] M) := +{ inv_fun := f.inv.to_fun, + left_inv := λ m, show (f.inv * f.val) m = m, + by erw f.inv_val; simp, + right_inv := λ m, show (f.val * f.inv) m = m, + by erw f.val_inv; simp, + ..f.val } + +/-- An equivalence from `M` to itself determines an invertible linear map. -/ +def of_linear_equiv (f : (M ≃ₗ[R] M)) : general_linear_group R M := +{ val := f, + inv := (f.symm : M →ₗ[R] M), + val_inv := linear_map.ext $ λ _, f.apply_symm_apply _, + inv_val := linear_map.ext $ λ _, f.symm_apply_apply _ } + +variables (R M) + +/-- The general linear group on `R` and `M` is multiplicatively equivalent to the type of linear +equivalences between `M` and itself. -/ +def general_linear_equiv : general_linear_group R M ≃* (M ≃ₗ[R] M) := +{ to_fun := to_linear_equiv, + inv_fun := of_linear_equiv, + left_inv := λ f, by { ext, refl }, + right_inv := λ f, by { ext, refl }, + map_mul' := λ x y, by {ext, refl} } + +@[simp] lemma general_linear_equiv_to_linear_map (f : general_linear_group R M) : + (general_linear_equiv R M f : M →ₗ[R] M) = f := +by {ext, refl} + +@[simp] lemma coe_fn_general_linear_equiv (f : general_linear_group R M) : + ⇑(general_linear_equiv R M f) = (f : M → M) := +rfl + +end general_linear_group + +end linear_map diff --git a/src/linear_algebra/matrix/general_linear_group.lean b/src/linear_algebra/matrix/general_linear_group.lean index 9a306aad8ad3b..5a6c6f6e35051 100644 --- a/src/linear_algebra/matrix/general_linear_group.lean +++ b/src/linear_algebra/matrix/general_linear_group.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Chris Birkbeck. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Birkbeck -/ +import linear_algebra.general_linear_group import linear_algebra.matrix.nonsingular_inverse import linear_algebra.matrix.special_linear_group diff --git a/src/linear_algebra/matrix/special_linear_group.lean b/src/linear_algebra/matrix/special_linear_group.lean index d371eab1cda05..b47f8eb80adfb 100644 --- a/src/linear_algebra/matrix/special_linear_group.lean +++ b/src/linear_algebra/matrix/special_linear_group.lean @@ -3,6 +3,7 @@ Copyright (c) 2020 Anne Baanen. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Anne Baanen -/ +import linear_algebra.general_linear_group import linear_algebra.matrix.adjugate import linear_algebra.matrix.to_lin diff --git a/src/linear_algebra/unitary_group.lean b/src/linear_algebra/unitary_group.lean index e27a84a586643..9a72696fc3684 100644 --- a/src/linear_algebra/unitary_group.lean +++ b/src/linear_algebra/unitary_group.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Shing Tak Lam. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Shing Tak Lam -/ +import linear_algebra.general_linear_group import linear_algebra.matrix.to_lin import linear_algebra.matrix.nonsingular_inverse import algebra.star.unitary diff --git a/src/number_theory/modular.lean b/src/number_theory/modular.lean index 3a4f23d563c54..168db9e4104a7 100644 --- a/src/number_theory/modular.lean +++ b/src/number_theory/modular.lean @@ -6,6 +6,7 @@ Authors: Alex Kontorovich, Heather Macbeth, Marc Masdeu import analysis.complex.upper_half_plane.basic import analysis.normed_space.finite_dimension +import linear_algebra.general_linear_group import linear_algebra.matrix.general_linear_group /-! From 84dc0bd6619acaea625086d6f53cb35cdd554219 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Fri, 3 Feb 2023 03:01:13 +0000 Subject: [PATCH 0416/1291] chore(topology/order/basic): split (#18363) Move material that requires topological groups to a new file to reduce transitive imports. --- .../box_integral/partition/basic.lean | 1 + src/analysis/convex/strict.lean | 2 +- src/topology/algebra/monoid.lean | 2 + src/topology/algebra/order/field.lean | 2 +- src/topology/algebra/order/floor.lean | 4 +- src/topology/algebra/order/group.lean | 80 +++++++++++ src/topology/algebra/order/liminf_limsup.lean | 2 + .../algebra/order/monotone_continuity.lean | 2 +- src/topology/continuous_function/ordered.lean | 1 + src/topology/order/basic.lean | 130 +++++------------- src/topology/order/lower_topology.lean | 1 + 11 files changed, 123 insertions(+), 104 deletions(-) create mode 100644 src/topology/algebra/order/group.lean diff --git a/src/analysis/box_integral/partition/basic.lean b/src/analysis/box_integral/partition/basic.lean index e9ff3fb379ac2..81704d3638187 100644 --- a/src/analysis/box_integral/partition/basic.lean +++ b/src/analysis/box_integral/partition/basic.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ +import algebra.big_operators.option import analysis.box_integral.box.basic /-! diff --git a/src/analysis/convex/strict.lean b/src/analysis/convex/strict.lean index f444b82976233..df09dbab49398 100644 --- a/src/analysis/convex/strict.lean +++ b/src/analysis/convex/strict.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ import analysis.convex.basic -import topology.order.basic +import topology.algebra.order.group /-! # Strictly convex sets diff --git a/src/topology/algebra/monoid.lean b/src/topology/algebra/monoid.lean index 3cccb7d1c337b..c430df3cdbc22 100644 --- a/src/topology/algebra/monoid.lean +++ b/src/topology/algebra/monoid.lean @@ -49,6 +49,8 @@ section has_continuous_mul variables [topological_space M] [has_mul M] [has_continuous_mul M] +@[to_additive] instance : has_continuous_mul Mᵒᵈ := ‹has_continuous_mul M› + @[to_additive] lemma continuous_mul : continuous (λp:M×M, p.1 * p.2) := has_continuous_mul.continuous_mul diff --git a/src/topology/algebra/order/field.lean b/src/topology/algebra/order/field.lean index e6802847b5b39..54a74468a91da 100644 --- a/src/topology/algebra/order/field.lean +++ b/src/topology/algebra/order/field.lean @@ -5,7 +5,7 @@ Authors: Benjamin Davidson, Devon Tuma, Eric Rodriguez, Oliver Nash -/ import tactic.positivity -import topology.order.basic +import topology.algebra.order.group import topology.algebra.field /-! diff --git a/src/topology/algebra/order/floor.lean b/src/topology/algebra/order/floor.lean index 9b7ccd9ab235d..2572bca7e14f4 100644 --- a/src/topology/algebra/order/floor.lean +++ b/src/topology/algebra/order/floor.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Anatole Dedecker -/ import algebra.order.floor -import topology.order.basic +import topology.algebra.order.group /-! # Topological facts about `int.floor`, `int.ceil` and `int.fract` @@ -156,7 +156,7 @@ tendsto_nhds_within_of_tendsto_nhds_of_eventually_within _ local notation `I` := (Icc 0 1 : set α) -variables [order_topology α] [topological_add_group α] [topological_space β] [topological_space γ] +variables [order_topology α] [topological_space β] [topological_space γ] /-- Do not use this, use `continuous_on.comp_fract` instead. -/ lemma continuous_on.comp_fract' {f : β → α → γ} diff --git a/src/topology/algebra/order/group.lean b/src/topology/algebra/order/group.lean new file mode 100644 index 0000000000000..a48d17fa14791 --- /dev/null +++ b/src/topology/algebra/order/group.lean @@ -0,0 +1,80 @@ +/- +Copyright (c) 2020 Yury Kudryashov. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yury Kudryashov +-/ +import topology.order.basic +import topology.algebra.group.basic + +/-! +# Topology on a linear ordered additive commutative group + +In this file we prove that a linear ordered additive commutative group with order topology is a +topological group. We also prove continuity of `abs : G → G` and provide convenience lemmas like +`continuous_at.abs`. +-/ + +open set filter +open_locale topology filter + +variables {α G : Type*} [topological_space G] [linear_ordered_add_comm_group G] [order_topology G] +variables {l : filter α} {f g : α → G} + +@[priority 100] -- see Note [lower instance priority] +instance linear_ordered_add_comm_group.topological_add_group : topological_add_group G := +{ continuous_add := + begin + refine continuous_iff_continuous_at.2 _, + rintro ⟨a, b⟩, + refine linear_ordered_add_comm_group.tendsto_nhds.2 (λ ε ε0, _), + rcases dense_or_discrete 0 ε with (⟨δ, δ0, δε⟩|⟨h₁, h₂⟩), + { -- If there exists `δ ∈ (0, ε)`, then we choose `δ`-nhd of `a` and `(ε-δ)`-nhd of `b` + filter_upwards [(eventually_abs_sub_lt a δ0).prod_nhds + (eventually_abs_sub_lt b (sub_pos.2 δε))], + rintros ⟨x, y⟩ ⟨hx : |x - a| < δ, hy : |y - b| < ε - δ⟩, + rw [add_sub_add_comm], + calc |x - a + (y - b)| ≤ |x - a| + |y - b| : abs_add _ _ + ... < δ + (ε - δ) : add_lt_add hx hy + ... = ε : add_sub_cancel'_right _ _ }, + { -- Otherwise `ε`-nhd of each point `a` is `{a}` + have hε : ∀ {x y}, |x - y| < ε → x = y, + { intros x y h, + simpa [sub_eq_zero] using h₂ _ h }, + filter_upwards [(eventually_abs_sub_lt a ε0).prod_nhds (eventually_abs_sub_lt b ε0)], + rintros ⟨x, y⟩ ⟨hx : |x - a| < ε, hy : |y - b| < ε⟩, + simpa [hε hx, hε hy] } + end, + continuous_neg := continuous_iff_continuous_at.2 $ λ a, + linear_ordered_add_comm_group.tendsto_nhds.2 $ λ ε ε0, + (eventually_abs_sub_lt a ε0).mono $ λ x hx, by rwa [neg_sub_neg, abs_sub_comm] } + +@[continuity] +lemma continuous_abs : continuous (abs : G → G) := continuous_id.max continuous_neg + +protected lemma filter.tendsto.abs {a : G} (h : tendsto f l (𝓝 a)) : + tendsto (λ x, |f x|) l (𝓝 (|a|)) := +(continuous_abs.tendsto _).comp h + +lemma tendsto_zero_iff_abs_tendsto_zero (f : α → G) : + tendsto f l (𝓝 0) ↔ tendsto (abs ∘ f) l (𝓝 0) := +begin + refine ⟨λ h, (abs_zero : |(0 : G)| = 0) ▸ h.abs, λ h, _⟩, + have : tendsto (λ a, -|f a|) l (𝓝 0) := (neg_zero : -(0 : G) = 0) ▸ h.neg, + exact tendsto_of_tendsto_of_tendsto_of_le_of_le this h + (λ x, neg_abs_le_self $ f x) (λ x, le_abs_self $ f x), +end + +variables [topological_space α] {a : α} {s : set α} + +protected lemma continuous.abs (h : continuous f) : continuous (λ x, |f x|) := continuous_abs.comp h + +protected lemma continuous_at.abs (h : continuous_at f a) : continuous_at (λ x, |f x|) a := h.abs + +protected lemma continuous_within_at.abs (h : continuous_within_at f s a) : + continuous_within_at (λ x, |f x|) s a := h.abs + +protected lemma continuous_on.abs (h : continuous_on f s) : continuous_on (λ x, |f x|) s := +λ x hx, (h x hx).abs + +lemma tendsto_abs_nhds_within_zero : tendsto (abs : G → G) (𝓝[≠] 0) (𝓝[>] 0) := +(continuous_abs.tendsto' (0 : G) 0 abs_zero).inf $ tendsto_principal_principal.2 $ λ x, abs_pos.2 diff --git a/src/topology/algebra/order/liminf_limsup.lean b/src/topology/algebra/order/liminf_limsup.lean index ef1bd9c837074..98c5b37d1a312 100644 --- a/src/topology/algebra/order/liminf_limsup.lean +++ b/src/topology/algebra/order/liminf_limsup.lean @@ -4,6 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro, Yury Kudryashov -/ import algebra.big_operators.intervals +import algebra.big_operators.order +import algebra.indicator_function import order.liminf_limsup import order.filter.archimedean import topology.order.basic diff --git a/src/topology/algebra/order/monotone_continuity.lean b/src/topology/algebra/order/monotone_continuity.lean index 2b18d6f919796..1b6e22bd0e399 100644 --- a/src/topology/algebra/order/monotone_continuity.lean +++ b/src/topology/algebra/order/monotone_continuity.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury G. Kudryashov, Heather Macbeth -/ import topology.order.basic -import topology.algebra.order.left_right +import topology.homeomorph /-! # Continuity of monotone functions diff --git a/src/topology/continuous_function/ordered.lean b/src/topology/continuous_function/ordered.lean index ce6e88b9f1c1a..e7073aa0bf720 100644 --- a/src/topology/continuous_function/ordered.lean +++ b/src/topology/continuous_function/ordered.lean @@ -5,6 +5,7 @@ Authors: Scott Morrison, Shing Tak Lam -/ import topology.algebra.order.proj_Icc +import topology.algebra.order.group import topology.continuous_function.basic /-! diff --git a/src/topology/order/basic.lean b/src/topology/order/basic.lean index be90cee93ee77..d7a7813ed49f0 100644 --- a/src/topology/order/basic.lean +++ b/src/topology/order/basic.lean @@ -6,7 +6,7 @@ Authors: Johannes Hölzl, Mario Carneiro, Yury Kudryashov import data.set.intervals.pi import data.set.pointwise.interval import order.filter.interval -import topology.algebra.group.basic +import topology.support import topology.algebra.order.left_right /-! @@ -94,9 +94,6 @@ instance [topological_space α] [h : first_countable_topology α] : first_counta 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 αᵒᵈ := h - lemma dense.order_dual [topological_space α] {s : set α} (hs : dense s) : dense (order_dual.of_dual ⁻¹' s) := hs @@ -1608,6 +1605,7 @@ end order_topology end linear_order section linear_ordered_add_comm_group + variables [topological_space α] [linear_ordered_add_comm_group α] [order_topology α] variables {l : filter β} {f g : β → α} @@ -1643,50 +1641,35 @@ lemma eventually_abs_sub_lt (a : α) {ε : α} (hε : 0 < ε) : ∀ᶠ x in 𝓝 (nhds_eq_infi_abs_sub a).symm ▸ mem_infi_of_mem ε (mem_infi_of_mem hε $ by simp only [abs_sub_comm, mem_principal_self]) -@[priority 100] -- see Note [lower instance priority] -instance linear_ordered_add_comm_group.topological_add_group : topological_add_group α := -{ continuous_add := - begin - refine continuous_iff_continuous_at.2 _, - rintro ⟨a, b⟩, - refine linear_ordered_add_comm_group.tendsto_nhds.2 (λ ε ε0, _), - rcases dense_or_discrete 0 ε with (⟨δ, δ0, δε⟩|⟨h₁, h₂⟩), - { -- If there exists `δ ∈ (0, ε)`, then we choose `δ`-nhd of `a` and `(ε-δ)`-nhd of `b` - filter_upwards [(eventually_abs_sub_lt a δ0).prod_nhds - (eventually_abs_sub_lt b (sub_pos.2 δε))], - rintros ⟨x, y⟩ ⟨hx : |x - a| < δ, hy : |y - b| < ε - δ⟩, - rw [add_sub_add_comm], - calc |x - a + (y - b)| ≤ |x - a| + |y - b| : abs_add _ _ - ... < δ + (ε - δ) : add_lt_add hx hy - ... = ε : add_sub_cancel'_right _ _ }, - { -- Otherwise `ε`-nhd of each point `a` is `{a}` - have hε : ∀ {x y}, |x - y| < ε → x = y, - { intros x y h, - simpa [sub_eq_zero] using h₂ _ h }, - filter_upwards [(eventually_abs_sub_lt a ε0).prod_nhds (eventually_abs_sub_lt b ε0)], - rintros ⟨x, y⟩ ⟨hx : |x - a| < ε, hy : |y - b| < ε⟩, - simpa [hε hx, hε hy] } - end, - continuous_neg := continuous_iff_continuous_at.2 $ λ a, - linear_ordered_add_comm_group.tendsto_nhds.2 $ λ ε ε0, - (eventually_abs_sub_lt a ε0).mono $ λ x hx, by rwa [neg_sub_neg, abs_sub_comm] } - -@[continuity] -lemma continuous_abs : continuous (abs : α → α) := continuous_id.max continuous_neg - -lemma filter.tendsto.abs {f : β → α} {a : α} {l : filter β} (h : tendsto f l (𝓝 a)) : - tendsto (λ x, |f x|) l (𝓝 (|a|)) := -(continuous_abs.tendsto _).comp h - -lemma tendsto_zero_iff_abs_tendsto_zero (f : β → α) {l : filter β} : - tendsto f l (𝓝 0) ↔ tendsto (abs ∘ f) l (𝓝 0) := +/-- In a linearly ordered additive commutative group with the order topology, if `f` tends to `C` +and `g` tends to `at_top` then `f + g` tends to `at_top`. -/ +lemma filter.tendsto.add_at_top {C : α} (hf : tendsto f l (𝓝 C)) (hg : tendsto g l at_top) : + tendsto (λ x, f x + g x) l at_top := begin - refine ⟨λ h, (abs_zero : |(0 : α)| = 0) ▸ h.abs, λ h, _⟩, - have : tendsto (λ a, -|f a|) l (𝓝 0) := (neg_zero : -(0 : α) = 0) ▸ h.neg, - exact tendsto_of_tendsto_of_tendsto_of_le_of_le this h - (λ x, neg_abs_le_self $ f x) (λ x, le_abs_self $ f x), + nontriviality α, + obtain ⟨C', hC'⟩ : ∃ C', C' < C := exists_lt C, + refine tendsto_at_top_add_left_of_le' _ C' _ hg, + exact (hf.eventually (lt_mem_nhds hC')).mono (λ x, le_of_lt) end +/-- In a linearly ordered additive commutative group with the order topology, if `f` tends to `C` +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 αᵒᵈ _ _ _ _ _ _ _ _ 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`. -/ +lemma filter.tendsto.at_top_add {C : α} (hf : tendsto f l at_top) (hg : tendsto g l (𝓝 C)) : + tendsto (λ x, f x + g x) l at_top := +by { conv in (_ + _) { rw add_comm }, exact hg.add_at_top hf } + +/-- In a linearly ordered additive commutative group with the order topology, if `f` tends to +`at_bot` and `g` tends to `C` then `f + g` tends to `at_bot`. -/ +lemma filter.tendsto.at_bot_add {C : α} (hf : tendsto f l at_bot) (hg : tendsto g l (𝓝 C)) : + tendsto (λ x, f x + g x) l at_bot := +by { conv in (_ + _) { rw add_comm }, exact hg.add_at_bot hf } + lemma nhds_basis_Ioo_pos [no_min_order α] [no_max_order α] (a : α) : (𝓝 a).has_basis (λ ε : α, (0 : α) < ε) (λ ε, Ioo (a-ε) (a+ε)) := ⟨begin @@ -1729,54 +1712,6 @@ lemma nhds_basis_Ioo_pos_of_pos [no_min_order α] [no_max_order α] (sub_le_sub_left (min_le_left i a) a) (add_le_add_left (min_le_left i a) a)) hit⟩, λ h, let ⟨i, hi, hit⟩ := h in ⟨i, hi.1, hit⟩ ⟩ ⟩ -section - -variables [topological_space β] {b : β} {a : α} {s : set β} - -lemma continuous.abs (h : continuous f) : continuous (λ x, |f x|) := continuous_abs.comp h - -lemma continuous_at.abs (h : continuous_at f b) : continuous_at (λ x, |f x|) b := h.abs - -lemma continuous_within_at.abs (h : continuous_within_at f s b) : - continuous_within_at (λ x, |f x|) s b := h.abs - -lemma continuous_on.abs (h : continuous_on f s) : continuous_on (λ x, |f x|) s := -λ x hx, (h x hx).abs - -lemma tendsto_abs_nhds_within_zero : tendsto (abs : α → α) (𝓝[≠] 0) (𝓝[>] 0) := -(continuous_abs.tendsto' (0 : α) 0 abs_zero).inf $ tendsto_principal_principal.2 $ λ x, abs_pos.2 - -end - -/-- In a linearly ordered additive commutative group with the order topology, if `f` tends to `C` -and `g` tends to `at_top` then `f + g` tends to `at_top`. -/ -lemma filter.tendsto.add_at_top {C : α} (hf : tendsto f l (𝓝 C)) (hg : tendsto g l at_top) : - tendsto (λ x, f x + g x) l at_top := -begin - nontriviality α, - obtain ⟨C', hC'⟩ : ∃ C', C' < C := exists_lt C, - refine tendsto_at_top_add_left_of_le' _ C' _ hg, - exact (hf.eventually (lt_mem_nhds hC')).mono (λ x, le_of_lt) -end - -/-- In a linearly ordered additive commutative group with the order topology, if `f` tends to `C` -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 αᵒᵈ _ _ _ _ _ _ _ _ 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`. -/ -lemma filter.tendsto.at_top_add {C : α} (hf : tendsto f l at_top) (hg : tendsto g l (𝓝 C)) : - tendsto (λ x, f x + g x) l at_top := -by { conv in (_ + _) { rw add_comm }, exact hg.add_at_top hf } - -/-- In a linearly ordered additive commutative group with the order topology, if `f` tends to -`at_bot` and `g` tends to `C` then `f + g` tends to `at_bot`. -/ -lemma filter.tendsto.at_bot_add {C : α} (hf : tendsto f l at_bot) (hg : tendsto g l (𝓝 C)) : - tendsto (λ x, f x + g x) l at_bot := -by { conv in (_ + _) { rw add_comm }, exact hg.add_at_bot hf } - end linear_ordered_add_comm_group lemma preimage_neg [add_group α] : preimage (has_neg.neg : α → α) = image (has_neg.neg : α → α) := @@ -2693,18 +2628,15 @@ section nhds_with_pos section linear_ordered_add_comm_group -variables [linear_ordered_add_comm_group α] [topological_space α] [order_topology α] +variables [linear_order α] [has_zero α] [topological_space α] [order_topology α] lemma eventually_nhds_within_pos_mem_Ioo {ε : α} (h : 0 < ε) : ∀ᶠ x in 𝓝[>] 0, x ∈ Ioo 0 ε := -begin - rw [eventually_iff, mem_nhds_within], - exact ⟨Ioo (-ε) ε, is_open_Ioo, by simp [h], λ x hx, ⟨hx.2, hx.1.2⟩⟩, -end +Ioo_mem_nhds_within_Ioi (left_mem_Ico.2 h) lemma eventually_nhds_within_pos_mem_Ioc {ε : α} (h : 0 < ε) : ∀ᶠ x in 𝓝[>] 0, x ∈ Ioc 0 ε := -(eventually_nhds_within_pos_mem_Ioo h).mono Ioo_subset_Ioc_self +Ioc_mem_nhds_within_Ioi (left_mem_Ico.2 h) end linear_ordered_add_comm_group diff --git a/src/topology/order/lower_topology.lean b/src/topology/order/lower_topology.lean index 496c12b200bb2..f5007759beb21 100644 --- a/src/topology/order/lower_topology.lean +++ b/src/topology/order/lower_topology.lean @@ -3,6 +3,7 @@ Copyright (c) 2023 Christopher Hoskin. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Christopher Hoskin -/ +import topology.homeomorph import topology.order.lattice import order.hom.complete_lattice From 0ebfdb71919ac6ca5d7fbc61a082fa2519556818 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Fri, 3 Feb 2023 08:20:41 +0000 Subject: [PATCH 0417/1291] chore(*): add mathlib4 synchronization comments (#18365) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.graded_mul_action` * `category_theory.sigma.basic` * `combinatorics.set_family.intersecting` * `combinatorics.set_family.lym` * `data.finsupp.alist` * `data.pfun` * `group_theory.presented_group` * `group_theory.quotient_group` * `group_theory.submonoid.inverses` * `order.ideal` * `set_theory.lists` * `topology.algebra.const_mul_action` * `topology.algebra.constructions` * `topology.connected` * `topology.dense_embedding` * `topology.homeomorph` * `topology.separation` * `topology.support` * `topology.uniform_space.basic` --- src/algebra/graded_mul_action.lean | 3 +++ src/category_theory/sigma/basic.lean | 3 +++ src/combinatorics/set_family/intersecting.lean | 3 +++ src/combinatorics/set_family/lym.lean | 3 +++ src/data/finsupp/alist.lean | 3 +++ src/data/pfun.lean | 3 +++ src/group_theory/presented_group.lean | 3 +++ src/group_theory/quotient_group.lean | 3 +++ src/group_theory/submonoid/inverses.lean | 3 +++ src/order/ideal.lean | 3 +++ src/set_theory/lists.lean | 3 +++ src/topology/algebra/const_mul_action.lean | 3 +++ src/topology/algebra/constructions.lean | 3 +++ src/topology/connected.lean | 3 +++ src/topology/dense_embedding.lean | 3 +++ src/topology/homeomorph.lean | 3 +++ src/topology/separation.lean | 3 +++ src/topology/support.lean | 3 +++ src/topology/uniform_space/basic.lean | 3 +++ 19 files changed, 57 insertions(+) diff --git a/src/algebra/graded_mul_action.lean b/src/algebra/graded_mul_action.lean index c70754ef199a2..041908d6f8e2d 100644 --- a/src/algebra/graded_mul_action.lean +++ b/src/algebra/graded_mul_action.lean @@ -8,6 +8,9 @@ import algebra.graded_monoid /-! # Additively-graded multiplicative action structures +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This module provides a set of heterogeneous typeclasses for defining a multiplicative structure over the sigma type `graded_monoid A` such that `(•) : A i → M j → M (i + j)`; that is to say, `A` has an additively-graded multiplicative action on `M`. The typeclasses are: diff --git a/src/category_theory/sigma/basic.lean b/src/category_theory/sigma/basic.lean index 8873ef5fca4f2..24e2fdd2b5c91 100644 --- a/src/category_theory/sigma/basic.lean +++ b/src/category_theory/sigma/basic.lean @@ -10,6 +10,9 @@ import category_theory.natural_isomorphism /-! # Disjoint union of categories +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define the category structure on a sigma-type (disjoint union) of categories. -/ diff --git a/src/combinatorics/set_family/intersecting.lean b/src/combinatorics/set_family/intersecting.lean index f90a3cb4e0781..b143e00a342b0 100644 --- a/src/combinatorics/set_family/intersecting.lean +++ b/src/combinatorics/set_family/intersecting.lean @@ -9,6 +9,9 @@ import order.upper_lower.basic /-! # Intersecting families +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines intersecting families and proves their basic properties. ## Main declarations diff --git a/src/combinatorics/set_family/lym.lean b/src/combinatorics/set_family/lym.lean index 750963e21c261..8d47f70ced332 100644 --- a/src/combinatorics/set_family/lym.lean +++ b/src/combinatorics/set_family/lym.lean @@ -12,6 +12,9 @@ import data.rat.order /-! # Lubell-Yamamoto-Meshalkin inequality and Sperner's theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves the local LYM and LYM inequalities as well as Sperner's theorem. ## Main declarations diff --git a/src/data/finsupp/alist.lean b/src/data/finsupp/alist.lean index 9ede0597051b9..fd8e54be5a929 100644 --- a/src/data/finsupp/alist.lean +++ b/src/data/finsupp/alist.lean @@ -9,6 +9,9 @@ import data.list.alist /-! # Connections between `finsupp` and `alist` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions * `finsupp.to_alist` diff --git a/src/data/pfun.lean b/src/data/pfun.lean index f698b62007161..588dcd7a2adbb 100644 --- a/src/data/pfun.lean +++ b/src/data/pfun.lean @@ -9,6 +9,9 @@ import data.rel /-! # Partial functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines partial functions. Partial functions are like functions, except they can also be "undefined" on some inputs. We define them as functions `α → part β`. diff --git a/src/group_theory/presented_group.lean b/src/group_theory/presented_group.lean index fcd3117a13e30..853596283152a 100644 --- a/src/group_theory/presented_group.lean +++ b/src/group_theory/presented_group.lean @@ -9,6 +9,9 @@ import group_theory.quotient_group /-! # Defining a group given by generators and relations +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given a subset `rels` of relations of the free group on a type `α`, this file constructs the group given by generators `x : α` and relations `r ∈ rels`. diff --git a/src/group_theory/quotient_group.lean b/src/group_theory/quotient_group.lean index 7d80d10bef0c1..f4145c16e1f07 100644 --- a/src/group_theory/quotient_group.lean +++ b/src/group_theory/quotient_group.lean @@ -13,6 +13,9 @@ import group_theory.subgroup.pointwise /-! # Quotients of groups by normal subgroups +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This files develops the basic theory of quotients of groups by normal subgroups. In particular it proves Noether's first and second isomorphism theorems. diff --git a/src/group_theory/submonoid/inverses.lean b/src/group_theory/submonoid/inverses.lean index 9cbbef13a63fc..c2ff17a5c081b 100644 --- a/src/group_theory/submonoid/inverses.lean +++ b/src/group_theory/submonoid/inverses.lean @@ -10,6 +10,9 @@ import group_theory.submonoid.pointwise # Submonoid of inverses +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given a submonoid `N` of a monoid `M`, we define the submonoid `N.left_inv` as the submonoid of left inverses of `N`. When `M` is commutative, we may define `from_comm_left_inv : N.left_inv →* N` since the inverses are unique. When `N ≤ is_unit.submonoid M`, this is precisely diff --git a/src/order/ideal.lean b/src/order/ideal.lean index f979b12efec09..f5a9afc2dca0a 100644 --- a/src/order/ideal.lean +++ b/src/order/ideal.lean @@ -10,6 +10,9 @@ import order.upper_lower.basic /-! # Order ideals, cofinal sets, and the Rasiowa–Sikorski lemma +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions Throughout this file, `P` is at least a preorder, but some sections require more diff --git a/src/set_theory/lists.lean b/src/set_theory/lists.lean index bf71507308f26..a9a067fbf96ee 100644 --- a/src/set_theory/lists.lean +++ b/src/set_theory/lists.lean @@ -8,6 +8,9 @@ import data.list.basic /-! # A computable model of ZFA without infinity +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define finite hereditary lists. This is useful for calculations in naive set theory. We distinguish two kinds of ZFA lists: diff --git a/src/topology/algebra/const_mul_action.lean b/src/topology/algebra/const_mul_action.lean index af3be61886696..cfd7af56a5712 100644 --- a/src/topology/algebra/const_mul_action.lean +++ b/src/topology/algebra/const_mul_action.lean @@ -12,6 +12,9 @@ import topology.support /-! # Monoid actions continuous in the second variable +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define class `has_continuous_const_smul`. We say `has_continuous_const_smul Γ T` if `Γ` acts on `T` and for each `γ`, the map `x ↦ γ • x` is continuous. (This differs from `has_continuous_smul`, which requires simultaneous continuity in both variables.) diff --git a/src/topology/algebra/constructions.lean b/src/topology/algebra/constructions.lean index 356f844454364..270d162e79400 100644 --- a/src/topology/algebra/constructions.lean +++ b/src/topology/algebra/constructions.lean @@ -8,6 +8,9 @@ import topology.homeomorph /-! # Topological space structure on the opposite monoid and on the units group +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define `topological_space` structure on `Mᵐᵒᵖ`, `Mᵃᵒᵖ`, `Mˣ`, and `add_units M`. This file does not import definitions of a topological monoid and/or a continuous multiplicative action, so we postpone the proofs of `has_continuous_mul Mᵐᵒᵖ` etc till we have these definitions. diff --git a/src/topology/connected.lean b/src/topology/connected.lean index 269c05f25224d..0dab6c95e6b2f 100644 --- a/src/topology/connected.lean +++ b/src/topology/connected.lean @@ -11,6 +11,9 @@ import tactic.congrm /-! # Connected subsets of topological spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define connected subsets of a topological spaces and various other properties and classes related to connectivity. diff --git a/src/topology/dense_embedding.lean b/src/topology/dense_embedding.lean index b2ddb767158b8..3e768af19cd82 100644 --- a/src/topology/dense_embedding.lean +++ b/src/topology/dense_embedding.lean @@ -9,6 +9,9 @@ import topology.bases /-! # Dense embeddings +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines three properties of functions: * `dense_range f` means `f` has dense image; diff --git a/src/topology/homeomorph.lean b/src/topology/homeomorph.lean index 9997b1fca2832..acab7ecacd254 100644 --- a/src/topology/homeomorph.lean +++ b/src/topology/homeomorph.lean @@ -10,6 +10,9 @@ import topology.support /-! # Homeomorphisms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines homeomorphisms between two topological spaces. They are bijections with both directions continuous. We denote homeomorphisms with the notation `≃ₜ`. diff --git a/src/topology/separation.lean b/src/topology/separation.lean index 1dafe29e01bc0..a4ef2b889c4ec 100644 --- a/src/topology/separation.lean +++ b/src/topology/separation.lean @@ -11,6 +11,9 @@ import topology.inseparable /-! # Separation properties of topological spaces. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the predicate `separated_nhds`, and common separation axioms (under the Kolmogorov classification). diff --git a/src/topology/support.lean b/src/topology/support.lean index 5e312403ab023..e52a6225b8039 100644 --- a/src/topology/support.lean +++ b/src/topology/support.lean @@ -9,6 +9,9 @@ import topology.separation /-! # The topological support of a function +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define the topological support of a function `f`, `tsupport f`, as the closure of the support of `f`. diff --git a/src/topology/uniform_space/basic.lean b/src/topology/uniform_space/basic.lean index c8a3c0692295d..178ffdb747e55 100644 --- a/src/topology/uniform_space/basic.lean +++ b/src/topology/uniform_space/basic.lean @@ -10,6 +10,9 @@ import topology.nhds_set /-! # Uniform spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Uniform spaces are a generalization of metric spaces and topological groups. Many concepts directly generalize to uniform spaces, e.g. From bed4f0529aa52e2b169d5c2359eef2849e306028 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 3 Feb 2023 11:16:34 +0000 Subject: [PATCH 0418/1291] chore(topology/algebra/infinite_sum): lemmas about pi.single (#18360) The naming is consistent with `finset.sum_pi_single`. --- src/topology/algebra/infinite_sum.lean | 8 ++++++++ 1 file changed, 8 insertions(+) diff --git a/src/topology/algebra/infinite_sum.lean b/src/topology/algebra/infinite_sum.lean index c8a486f123e4d..653f696bcd0f1 100644 --- a/src/topology/algebra/infinite_sum.lean +++ b/src/topology/algebra/infinite_sum.lean @@ -163,6 +163,10 @@ begin exact if_neg hb' end +lemma has_sum_pi_single [decidable_eq β] (b : β) (a : α) : + has_sum (pi.single b a) a := +show has_sum (λ x, pi.single b a x) a, by simpa only [pi.single_apply] using has_sum_ite_eq b a + lemma equiv.has_sum_iff (e : γ ≃ β) : has_sum (f ∘ e) a ↔ has_sum f a := e.injective.has_sum_iff $ by simp @@ -506,6 +510,10 @@ calc ∑' b' c', f b' c' = ∑' b', f b' c : tsum_congr $ λ b', tsum_eq_single ∑' b', (if b' = b then a else 0) = a := (has_sum_ite_eq b a).tsum_eq +@[simp] lemma tsum_pi_single [decidable_eq β] (b : β) (a : α) : + ∑' b', pi.single b a b' = a := +(has_sum_pi_single b a).tsum_eq + lemma tsum_dite_right (P : Prop) [decidable P] (x : β → ¬ P → α) : ∑' (b : β), (if h : P then (0 : α) else x b h) = if h : P then (0 : α) else ∑' (b : β), x b h := by by_cases hP : P; simp [hP] From 92ca63f0fb391a9ca5f22d2409a6080e786d99f7 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Fri, 3 Feb 2023 13:16:56 +0000 Subject: [PATCH 0419/1291] refactor(tactic/wlog): simplify and speed up `wlog` (#16495) Benefits: - The tactic is faster - The tactic is easier to port to Lean 4 Downside: - The tactic doesn't do any heavy-lifting for the user Zulip thread: https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/wlog/near/296996966 Co-authored-by: Yury G. Kudryashov Co-authored-by: Scott Morrison --- archive/imo/imo1988_q6.lean | 3 +- archive/imo/imo2006_q3.lean | 17 +- src/algebra/order/hom/ring.lean | 6 +- src/analysis/bounded_variation.lean | 16 +- .../box_integral/partition/filter.lean | 15 +- src/analysis/convex/function.lean | 14 +- src/analysis/normed_space/ray.lean | 12 +- src/analysis/special_functions/pow.lean | 20 +- src/combinatorics/composition.lean | 6 +- src/computability/DFA.lean | 3 +- src/computability/turing_machine.lean | 5 +- src/control/lawful_fix.lean | 3 +- src/data/complex/basic.lean | 16 +- src/data/fintype/card.lean | 3 +- src/data/nat/count.lean | 4 +- src/data/nat/fib.lean | 16 +- src/data/nat/nth.lean | 1 + src/data/set/enumerate.lean | 4 +- .../intervals/ord_connected_component.lean | 4 +- src/dynamics/ergodic/measure_preserving.lean | 11 +- src/field_theory/separable.lean | 6 +- src/field_theory/separable_degree.lean | 3 +- src/group_theory/order_of_element.lean | 3 +- src/group_theory/perm/cycle/basic.lean | 3 +- .../exterior_algebra/basic.lean | 10 +- .../constructions/borel_space.lean | 17 +- src/measure_theory/covering/besicovitch.lean | 28 +- .../covering/besicovitch_vector_space.lean | 10 +- .../integral/divergence_theorem.lean | 48 +-- .../integral/interval_integral.lean | 15 +- .../number_field/embeddings.lean | 8 +- src/number_theory/padics/padic_norm.lean | 7 +- src/order/omega_complete_partial_order.lean | 2 +- src/ring_theory/roots_of_unity.lean | 3 +- src/ring_theory/valuation/basic.lean | 16 +- src/set_theory/cardinal/divisibility.lean | 7 +- src/set_theory/game/nim.lean | 17 +- src/set_theory/surreal/dyadic.lean | 1 + src/tactic/wlog.lean | 285 ++++-------------- .../algebra/order/left_right_lim.lean | 16 +- src/topology/algebra/with_zero_topology.lean | 6 +- src/topology/connected.lean | 4 +- src/topology/instances/rat_lemmas.lean | 3 +- src/topology/metric_space/basic.lean | 10 +- .../metric_space/emetric_paracompact.lean | 5 +- src/topology/separation.lean | 14 +- src/topology/shrinking_lemma.lean | 3 +- test/wlog.lean | 135 ++------- 48 files changed, 309 insertions(+), 555 deletions(-) diff --git a/archive/imo/imo1988_q6.lean b/archive/imo/imo1988_q6.lean index 5db7f6694d4fb..24f0137c54364 100644 --- a/archive/imo/imo1988_q6.lean +++ b/archive/imo/imo1988_q6.lean @@ -70,7 +70,8 @@ lemma constant_descent_vieta_jumping (x y : ℕ) {claim : Prop} {H : ℕ → ℕ begin -- First of all, we may assume that x ≤ y. -- We justify this using H_symm. - wlog hxy : x ≤ y, swap, { rw H_symm at h₀, solve_by_elim }, + wlog hxy : x ≤ y, + { rw H_symm at h₀, apply this y x h₀ B C base _ _ _ _ _ _ (le_of_not_le hxy), assumption' }, -- In fact, we can easily deal with the case x = y. by_cases x_eq_y : x = y, {subst x_eq_y, exact H_diag h₀}, -- Hence we may assume that x < y. diff --git a/archive/imo/imo2006_q3.lean b/archive/imo/imo2006_q3.lean index b2c2b38fc99f6..c6506dbcda082 100644 --- a/archive/imo/imo2006_q3.lean +++ b/archive/imo/imo2006_q3.lean @@ -89,17 +89,16 @@ le_of_pow_le_pow _ (mul_nonneg (sqrt_nonneg _) (sq_nonneg _)) nat.succ_pos' $ theorem subst_proof₁ (x y z s : ℝ) (hxyz : x + y + z = 0) : |x * y * z * s| ≤ sqrt 2 / 32 * (x^2 + y^2 + z^2 + s^2)^2 := begin - wlog h' := mul_nonneg_of_three x y z using [x y z, y z x, z x y] tactic.skip, + wlog h' : 0 ≤ x * y generalizing x y z, swap, { rw [div_mul_eq_mul_div, le_div_iff' zero_lt_32], exact subst_wlog h' hxyz }, - { intro h, - rw [add_assoc, add_comm] at h, - rw [mul_assoc x, mul_comm x, add_assoc (x^2), add_comm (x^2)], - exact this h }, - { intro h, - rw [add_comm, ← add_assoc] at h, - rw [mul_comm _ z, ← mul_assoc, add_comm _ (z^2), ← add_assoc], - exact this h } + cases (mul_nonneg_of_three x y z).resolve_left h' with h h, + { specialize this y z x _ h, + { rw ← hxyz, ring, }, + { convert this using 2; ring } }, + { specialize this z x y _ h, + { rw ← hxyz, ring, }, + { convert this using 2; ring } }, end lemma lhs_identity (a b c : ℝ) : diff --git a/src/algebra/order/hom/ring.lean b/src/algebra/order/hom/ring.lean index 6cc404b68cf77..1c34a6c865df5 100644 --- a/src/algebra/order/hom/ring.lean +++ b/src/algebra/order/hom/ring.lean @@ -325,9 +325,9 @@ instance order_ring_hom.subsingleton [linear_ordered_field α] [linear_ordered_f subsingleton (α →+*o β) := ⟨λ f g, begin ext x, - by_contra' h, - wlog h : f x < g x using [f g, g f], - { exact ne.lt_or_lt h }, + by_contra' h' : f x ≠ g x, + wlog h : f x < g x, + { exact this g f x (ne.symm h') (h'.lt_or_lt.resolve_left h), }, obtain ⟨q, hf, hg⟩ := exists_rat_btwn h, rw ←map_rat_cast f at hf, rw ←map_rat_cast g at hg, diff --git a/src/analysis/bounded_variation.lean b/src/analysis/bounded_variation.lean index 7e74573ea7701..d1db929d718b9 100644 --- a/src/analysis/bounded_variation.lean +++ b/src/analysis/bounded_variation.lean @@ -181,10 +181,9 @@ lemma _root_.has_bounded_variation_on.has_locally_bounded_variation_on {f : α lemma edist_le (f : α → E) {s : set α} {x y : α} (hx : x ∈ s) (hy : y ∈ s) : edist (f x) (f y) ≤ evariation_on f s := begin - wlog hxy : x ≤ y := le_total x y using [x y, y x] tactic.skip, swap, - { assume hx hy, - rw edist_comm, - exact this hy hx }, + wlog hxy : x ≤ y, + { rw edist_comm, + exact this f hy hx (le_of_not_le hxy) }, let u : ℕ → α := λ n, if n = 0 then x else y, have hu : monotone u, { assume m n hmn, @@ -798,15 +797,14 @@ lemma edist_zero_of_eq_zero {a b : α} (ha : a ∈ s) (hb : b ∈ s) (h : variation_on_from_to f s a b = 0) : edist (f a) (f b) = 0 := begin - wlog h' : a ≤ b := le_total a b using [b a, a b] tactic.skip, + wlog h' : a ≤ b, + { rw edist_comm, + apply this hf hb ha _ (le_of_not_le h'), + rw [eq_neg_swap, h, neg_zero] }, { apply le_antisymm _ (zero_le _), rw [←ennreal.of_real_zero, ←h, eq_of_le f s h', ennreal.of_real_to_real (hf a b ha hb)], apply evariation_on.edist_le, exacts [⟨ha, ⟨le_rfl, h'⟩⟩, ⟨hb, ⟨h', le_rfl⟩⟩] }, - { assume ha hb hab, - rw edist_comm, - apply this hb ha, - rw [eq_neg_swap, hab, neg_zero] } end @[protected] diff --git a/src/analysis/box_integral/partition/filter.lean b/src/analysis/box_integral/partition/filter.lean index 7f33e1c02fede..d797111834b4a 100644 --- a/src/analysis/box_integral/partition/filter.lean +++ b/src/analysis/box_integral/partition/filter.lean @@ -333,14 +333,13 @@ lemma mem_base_set.exists_common_compl (h₁ : l.mem_base_set I c₁ r₁ π₁) ∃ π : prepartition I, π.Union = I \ π₁.Union ∧ (l.bDistortion → π.distortion ≤ c₁) ∧ (l.bDistortion → π.distortion ≤ c₂) := begin - wlog hc : c₁ ≤ c₂ := le_total c₁ c₂ using [c₁ c₂ r₁ r₂ π₁ π₂, c₂ c₁ r₂ r₁ π₂ π₁] tactic.skip, - { by_cases hD : (l.bDistortion : Prop), - { rcases h₁.4 hD with ⟨π, hπU, hπc⟩, - exact ⟨π, hπU, λ _, hπc, λ _, hπc.trans hc⟩ }, - { exact ⟨π₁.to_prepartition.compl, π₁.to_prepartition.Union_compl, - λ h, (hD h).elim, λ h, (hD h).elim⟩ } }, - { intros h₁ h₂ hU, - simpa [hU, and_comm] using this h₂ h₁ hU.symm } + wlog hc : c₁ ≤ c₂, + { simpa [hU, and_comm] using this h₂ h₁ hU.symm (le_of_not_le hc) }, + by_cases hD : (l.bDistortion : Prop), + { rcases h₁.4 hD with ⟨π, hπU, hπc⟩, + exact ⟨π, hπU, λ _, hπc, λ _, hπc.trans hc⟩ }, + { exact ⟨π₁.to_prepartition.compl, π₁.to_prepartition.Union_compl, + λ h, (hD h).elim, λ h, (hD h).elim⟩ } end protected lemma mem_base_set.union_compl_to_subordinate (hπ₁ : l.mem_base_set I c r₁ π₁) diff --git a/src/analysis/convex/function.lean b/src/analysis/convex/function.lean index 920791e909dd6..ef5948f6e871f 100644 --- a/src/analysis/convex/function.lean +++ b/src/analysis/convex/function.lean @@ -342,9 +342,10 @@ lemma linear_order.convex_on_of_lt (hs : convex 𝕜 s) f (a • x + b • y) ≤ a • f x + b • f y) : convex_on 𝕜 s f := begin refine convex_on_iff_pairwise_pos.2 ⟨hs, λ x hx y hy hxy a b ha hb hab, _⟩, - wlog h : x ≤ y using [x y a b, y x b a], - { exact le_total _ _ }, - exact hf hx hy (h.lt_of_ne hxy) ha hb hab, + wlog h : x < y, + { rw [add_comm (a • x), add_comm (a • f x)], rw add_comm at hab, + refine this hs hf y hy x hx hxy.symm b a hb ha hab (hxy.lt_or_lt.resolve_left h), }, + exact hf hx hy h ha hb hab, end /-- For a function on a convex set in a linearly ordered space (where the order and the algebraic @@ -365,9 +366,10 @@ lemma linear_order.strict_convex_on_of_lt (hs : convex 𝕜 s) f (a • x + b • y) < a • f x + b • f y) : strict_convex_on 𝕜 s f := begin refine ⟨hs, λ x hx y hy hxy a b ha hb hab, _⟩, - wlog h : x ≤ y using [x y a b, y x b a], - { exact le_total _ _ }, - exact hf hx hy (h.lt_of_ne hxy) ha hb hab, + wlog h : x < y, + { rw [add_comm (a • x), add_comm (a • f x)], rw add_comm at hab, + refine this hs hf y hy x hx hxy.symm b a hb ha hab (hxy.lt_or_lt.resolve_left h), }, + exact hf hx hy h ha hb hab, end /-- For a function on a convex set in a linearly ordered space (where the order and the algebraic diff --git a/src/analysis/normed_space/ray.lean b/src/analysis/normed_space/ray.lean index eb01ab23957e8..f2e63ce6ee08d 100644 --- a/src/analysis/normed_space/ray.lean +++ b/src/analysis/normed_space/ray.lean @@ -36,12 +36,12 @@ end lemma norm_sub (h : same_ray ℝ x y) : ‖x - y‖ = |‖x‖ - ‖y‖| := begin rcases h.exists_eq_smul with ⟨u, a, b, ha, hb, -, rfl, rfl⟩, - wlog hab : b ≤ a := le_total b a using [a b, b a] tactic.skip, - { rw ← sub_nonneg at hab, - rw [← sub_smul, norm_smul_of_nonneg hab, norm_smul_of_nonneg ha, - norm_smul_of_nonneg hb, ← sub_mul, abs_of_nonneg (mul_nonneg hab (norm_nonneg _))] }, - { intros ha hb hab, - rw [norm_sub_rev, this hb ha hab.symm, abs_sub_comm] } + wlog hab : b ≤ a, + { rw same_ray_comm at h, rw [norm_sub_rev, abs_sub_comm], + exact this u b a hb ha h (le_of_not_le hab), }, + rw ← sub_nonneg at hab, + rw [← sub_smul, norm_smul_of_nonneg hab, norm_smul_of_nonneg ha, + norm_smul_of_nonneg hb, ← sub_mul, abs_of_nonneg (mul_nonneg hab (norm_nonneg _))] end lemma norm_smul_eq (h : same_ray ℝ x y) : ‖x‖ • y = ‖y‖ • x := diff --git a/src/analysis/special_functions/pow.lean b/src/analysis/special_functions/pow.lean index c7cb1b7f2eaf0..3aec5d760b620 100644 --- a/src/analysis/special_functions/pow.lean +++ b/src/analysis/special_functions/pow.lean @@ -1783,16 +1783,16 @@ lemma mul_rpow_eq_ite (x y : ℝ≥0∞) (z : ℝ) : begin rcases eq_or_ne z 0 with rfl|hz, { simp }, replace hz := hz.lt_or_lt, - wlog hxy : x ≤ y := le_total x y using [x y, y x] tactic.skip, - { rcases eq_or_ne x 0 with rfl|hx0, - { induction y using with_top.rec_top_coe; cases hz with hz hz; simp [*, hz.not_lt] }, - rcases eq_or_ne y 0 with rfl|hy0, { exact (hx0 (bot_unique hxy)).elim }, - induction x using with_top.rec_top_coe, { cases hz with hz hz; simp [hz, top_unique hxy] }, - induction y using with_top.rec_top_coe, { cases hz with hz hz; simp * }, - simp only [*, false_and, and_false, false_or, if_false], - norm_cast at *, - rw [coe_rpow_of_ne_zero (mul_ne_zero hx0 hy0), nnreal.mul_rpow] }, - { convert this using 2; simp only [mul_comm, and_comm, or_comm] } + wlog hxy : x ≤ y, + { convert this y x z hz (le_of_not_le hxy) using 2; simp only [mul_comm, and_comm, or_comm], }, + rcases eq_or_ne x 0 with rfl|hx0, + { induction y using with_top.rec_top_coe; cases hz with hz hz; simp [*, hz.not_lt] }, + rcases eq_or_ne y 0 with rfl|hy0, { exact (hx0 (bot_unique hxy)).elim }, + induction x using with_top.rec_top_coe, { cases hz with hz hz; simp [hz, top_unique hxy] }, + induction y using with_top.rec_top_coe, { cases hz with hz hz; simp * }, + simp only [*, false_and, and_false, false_or, if_false], + norm_cast at *, + rw [coe_rpow_of_ne_zero (mul_ne_zero hx0 hy0), nnreal.mul_rpow] end lemma mul_rpow_of_ne_top {x y : ℝ≥0∞} (hx : x ≠ ⊤) (hy : y ≠ ⊤) (z : ℝ) : diff --git a/src/combinatorics/composition.lean b/src/combinatorics/composition.lean index c7fad7f0386ec..235638b9f8347 100644 --- a/src/combinatorics/composition.lean +++ b/src/combinatorics/composition.lean @@ -341,14 +341,12 @@ lemma disjoint_range {i₁ i₂ : fin c.length} (h : i₁ ≠ i₂) : disjoint (set.range (c.embedding i₁)) (set.range (c.embedding i₂)) := begin classical, - wlog h' : i₁ ≤ i₂ using i₁ i₂, - swap, exact (this h.symm).symm, + wlog h' : i₁ < i₂, { exact (this c h.symm (h.lt_or_lt.resolve_left h')).symm }, by_contradiction d, obtain ⟨x, hx₁, hx₂⟩ : ∃ x : fin n, (x ∈ set.range (c.embedding i₁) ∧ x ∈ set.range (c.embedding i₂)) := set.not_disjoint_iff.1 d, - have : i₁ < i₂ := lt_of_le_of_ne h' h, - have A : (i₁ : ℕ).succ ≤ i₂ := nat.succ_le_of_lt this, + have A : (i₁ : ℕ).succ ≤ i₂ := nat.succ_le_of_lt h', apply lt_irrefl (x : ℕ), calc (x : ℕ) < c.size_up_to (i₁ : ℕ).succ : (c.mem_range_embedding_iff.1 hx₁).2 ... ≤ c.size_up_to (i₂ : ℕ) : monotone_sum_take _ A diff --git a/src/computability/DFA.lean b/src/computability/DFA.lean index 0c815938b9a00..948b28ad9ce6c 100644 --- a/src/computability/DFA.lean +++ b/src/computability/DFA.lean @@ -77,8 +77,7 @@ lemma eval_from_split [fintype σ] {x : list α} {s t : σ} (hlen : fintype.card begin obtain ⟨n, m, hneq, heq⟩ := fintype.exists_ne_map_eq_of_card_lt (λ n : fin (fintype.card σ + 1), M.eval_from s (x.take n)) (by norm_num), - wlog hle : (n : ℕ) ≤ m using n m, - have hlt : (n : ℕ) < m := (ne.le_iff_lt hneq).mp hle, + wlog hle : (n : ℕ) ≤ m, { exact this hlen hx _ _ hneq.symm heq.symm (le_of_not_le hle), }, have hm : (m : ℕ) ≤ fintype.card σ := fin.is_le m, dsimp at heq, diff --git a/src/computability/turing_machine.lean b/src/computability/turing_machine.lean index 7924f8ed6a3a4..4bb8e26ed37fc 100644 --- a/src/computability/turing_machine.lean +++ b/src/computability/turing_machine.lean @@ -268,8 +268,9 @@ end (∀ i, L₁.nth i = L₂.nth i) → L₁ = L₂ := list_blank.induction_on L₁ $ λ l₁, list_blank.induction_on L₂ $ λ l₂ H, begin - wlog h : l₁.length ≤ l₂.length using l₁ l₂, - swap, { exact (this $ λ i, (H i).symm).symm }, + wlog h : l₁.length ≤ l₂.length, + { cases le_total l₁.length l₂.length; [skip, symmetry]; apply_assumption; try {assumption}, + intro, rw H }, refine quotient.sound' (or.inl ⟨l₂.length - l₁.length, _⟩), refine list.ext_le _ (λ i h h₂, eq.symm _), { simp only [add_tsub_cancel_of_le h, list.length_append, list.length_replicate] }, diff --git a/src/control/lawful_fix.lean b/src/control/lawful_fix.lean index 64df81a2c9ad0..817a2e954ef45 100644 --- a/src/control/lawful_fix.lean +++ b/src/control/lawful_fix.lean @@ -74,7 +74,8 @@ begin suffices : y = b, subst this, exact h₁, cases hh with i hh, revert h₁, generalize : (succ (nat.find h₀)) = j, intro, - wlog : i ≤ j := le_total i j using [i j b y,j i y b], + wlog case : i ≤ j, + { cases le_total i j with H H; [skip, symmetry]; apply_assumption; assumption }, replace hh := approx_mono f case _ _ hh, apply part.mem_unique h₁ hh }, { simp only [fix_def' ⇑f h₀, not_exists, false_iff, not_mem_none], diff --git a/src/data/complex/basic.lean b/src/data/complex/basic.lean index 91a575a64b484..dfbe4449da6fb 100644 --- a/src/data/complex/basic.lean +++ b/src/data/complex/basic.lean @@ -582,15 +582,13 @@ by simpa [re_add_im] using abs.add_le z.re (z.im * I) lemma abs_le_sqrt_two_mul_max (z : ℂ) : abs z ≤ real.sqrt 2 * max (|z.re|) (|z.im|) := begin cases z with x y, - simp only [abs, norm_sq_mk, ← sq], - wlog hle : |x| ≤ |y| := le_total (|x|) (|y|) using [x y, y x] tactic.skip, - { simp only [absolute_value.coe_mk, mul_hom.coe_mk, norm_sq_mk, ←sq], - calc real.sqrt (x ^ 2 + y ^ 2) ≤ real.sqrt (y ^ 2 + y ^ 2) : - real.sqrt_le_sqrt (add_le_add_right (sq_le_sq.2 hle) _) - ... = real.sqrt 2 * max (|x|) (|y|) : - by rw [max_eq_right hle, ← two_mul, real.sqrt_mul two_pos.le, real.sqrt_sq_eq_abs] }, - { dsimp, - rwa [add_comm, max_comm] } + simp only [abs_apply, norm_sq_mk, ← sq], + wlog hle : |x| ≤ |y|, + { rw [add_comm, max_comm], exact this _ _ (le_of_not_le hle), }, + calc real.sqrt (x ^ 2 + y ^ 2) ≤ real.sqrt (y ^ 2 + y ^ 2) : + real.sqrt_le_sqrt (add_le_add_right (sq_le_sq.2 hle) _) + ... = real.sqrt 2 * max (|x|) (|y|) : + by rw [max_eq_right hle, ← two_mul, real.sqrt_mul two_pos.le, real.sqrt_sq_eq_abs], end lemma abs_re_div_abs_le_one (z : ℂ) : |z.re / z.abs| ≤ 1 := diff --git a/src/data/fintype/card.lean b/src/data/fintype/card.lean index ab990add5a863..43d7bf3b7f0f5 100644 --- a/src/data/fintype/card.lean +++ b/src/data/fintype/card.lean @@ -864,7 +864,8 @@ private lemma nat_embedding_aux_injective (α : Type*) [infinite α] : begin rintro m n h, letI := classical.dec_eq α, - wlog hmlen : m ≤ n using m n, + wlog hmlen : m ≤ n generalizing m n, + { exact (this h.symm $ le_of_not_le hmlen).symm }, by_contradiction hmn, have hmn : m < n, from lt_of_le_of_ne hmlen hmn, refine (classical.some_spec (exists_not_mem_finset diff --git a/src/data/nat/count.lean b/src/data/nat/count.lean index 94ccd876517e2..a35fd9efa9ee8 100644 --- a/src/data/nat/count.lean +++ b/src/data/nat/count.lean @@ -102,9 +102,9 @@ lemma count_strict_mono {m n : ℕ} (hm : p m) (hmn : m < n) : count p m < count lemma count_injective {m n : ℕ} (hm : p m) (hn : p n) (heq : count p m = count p n) : m = n := begin - by_contra, + by_contra' h : m ≠ n, wlog hmn : m < n, - { exact ne.lt_or_lt h }, + { exact this hn hm heq.symm h.symm (h.lt_or_lt.resolve_left hmn) }, { simpa [heq] using count_strict_mono hm hmn } end diff --git a/src/data/nat/fib.lean b/src/data/nat/fib.lean index a56d0fa9c1c7c..9c562beb94455 100644 --- a/src/data/nat/fib.lean +++ b/src/data/nat/fib.lean @@ -222,15 +222,13 @@ lemma gcd_fib_add_mul_self (m n : ℕ) : ∀ k, gcd (fib m) (fib (n + k * m)) = see https://proofwiki.org/wiki/GCD_of_Fibonacci_Numbers -/ lemma fib_gcd (m n : ℕ) : fib (gcd m n) = gcd (fib m) (fib n) := begin - wlog h : m ≤ n using [n m, m n], - exact le_total m n, - { apply gcd.induction m n, - { simp }, - intros m n mpos h, - rw ← gcd_rec m n at h, - conv_rhs { rw ← mod_add_div' n m }, - rwa [gcd_fib_add_mul_self m (n % m) (n / m), gcd_comm (fib m) _] }, - rwa [gcd_comm, gcd_comm (fib m)] + wlog h : m ≤ n, { simpa only [gcd_comm] using this _ _ (le_of_not_le h) }, + apply gcd.induction m n, + { simp }, + intros m n mpos h, + rw ← gcd_rec m n at h, + conv_rhs { rw ← mod_add_div' n m }, + rwa [gcd_fib_add_mul_self m (n % m) (n / m), gcd_comm (fib m) _] end lemma fib_dvd (m n : ℕ) (h : m ∣ n) : fib m ∣ fib n := diff --git a/src/data/nat/nth.lean b/src/data/nat/nth.lean index c1acd39cc0617..9f4da094252c4 100644 --- a/src/data/nat/nth.lean +++ b/src/data/nat/nth.lean @@ -150,6 +150,7 @@ lemma nth_injective_of_infinite (hp : (set_of p).infinite) : function.injective begin intros m n h, wlog h' : m ≤ n, + { exact (this p hp h.symm (le_of_not_le h')).symm }, rw le_iff_lt_or_eq at h', obtain (h' | rfl) := h', { simpa [h] using nth_strict_mono p hp h' }, diff --git a/src/data/set/enumerate.lean b/src/data/set/enumerate.lean index 8343cdeda7c1e..78fe97525dd95 100644 --- a/src/data/set/enumerate.lean +++ b/src/data/set/enumerate.lean @@ -3,8 +3,9 @@ Copyright (c) 2017 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl -/ -import tactic.wlog +import data.set.lattice import data.nat.order.basic +import tactic.wlog /-! # Set enumeration @@ -72,6 +73,7 @@ lemma enumerate_inj {n₁ n₂ : ℕ} {a : α} {s : set α} (h_sel : ∀ s a, se (h₁ : enumerate s n₁ = some a) (h₂ : enumerate s n₂ = some a) : n₁ = n₂ := begin wlog hn : n₁ ≤ n₂, + { cases le_total n₁ n₂ with H H; [skip, symmetry]; apply_assumption; assumption }, { rcases nat.le.dest hn with ⟨m, rfl⟩, clear hn, induction n₁ generalizing s, case nat.zero diff --git a/src/data/set/intervals/ord_connected_component.lean b/src/data/set/intervals/ord_connected_component.lean index eff4c832d07d6..e425e8a41bc87 100644 --- a/src/data/set/intervals/ord_connected_component.lean +++ b/src/data/set/intervals/ord_connected_component.lean @@ -165,8 +165,8 @@ begin rcases mem_Union₂.1 hx₁ with ⟨a, has, ha⟩, clear hx₁, rcases mem_Union₂.1 hx₂ with ⟨b, hbt, hb⟩, clear hx₂, rw [mem_ord_connected_component, subset_inter_iff] at ha hb, - wlog hab : a ≤ b := le_total a b using [a b s t, b a t s] tactic.skip, - rotate, from λ h₁ h₂ h₃ h₄, this h₂ h₁ h₄ h₃, + wlog hab : a ≤ b, + { exact this b hbt a has ha hb (le_of_not_le hab) }, cases ha with ha ha', cases hb with hb hb', have hsub : [a, b] ⊆ (ord_separating_set s t).ord_connected_sectionᶜ, { rw [ord_separating_set_comm, uIcc_comm] at hb', diff --git a/src/dynamics/ergodic/measure_preserving.lean b/src/dynamics/ergodic/measure_preserving.lean index 751332d237ae3..8a11e4d7efa6d 100644 --- a/src/dynamics/ergodic/measure_preserving.lean +++ b/src/dynamics/ergodic/measure_preserving.lean @@ -137,12 +137,11 @@ begin by simpa only [B, nsmul_eq_mul, finset.sum_const, finset.card_range], rcases exists_nonempty_inter_of_measure_univ_lt_sum_measure μ (λ m hm, A m) this with ⟨i, hi, j, hj, hij, x, hxi, hxj⟩, - -- without `tactic.skip` Lean closes the extra goal but it takes a long time; not sure why - wlog hlt : i < j := hij.lt_or_lt using [i j, j i] tactic.skip, - { simp only [set.mem_preimage, finset.mem_range] at hi hj hxi hxj, - refine ⟨f^[i] x, hxi, j - i, ⟨tsub_pos_of_lt hlt, lt_of_le_of_lt (j.sub_le i) hj⟩, _⟩, - rwa [← iterate_add_apply, tsub_add_cancel_of_le hlt.le] }, - { exact λ hi hj hij hxi hxj, this hj hi hij.symm hxj hxi } + wlog hlt : i < j generalizing i j, + { exact this j hj i hi hij.symm hxj hxi (hij.lt_or_lt.resolve_left hlt) }, + simp only [set.mem_preimage, finset.mem_range] at hi hj hxi hxj, + refine ⟨f^[i] x, hxi, j - i, ⟨tsub_pos_of_lt hlt, lt_of_le_of_lt (j.sub_le i) hj⟩, _⟩, + rwa [← iterate_add_apply, tsub_add_cancel_of_le hlt.le] end /-- A self-map preserving a finite measure is conservative: if `μ s ≠ 0`, then at least one point diff --git a/src/field_theory/separable.lean b/src/field_theory/separable.lean index ffbe6f49b9320..af29f7b6c2003 100644 --- a/src/field_theory/separable.lean +++ b/src/field_theory/separable.lean @@ -335,7 +335,9 @@ theorem unique_separable_of_irreducible {f : F[X]} (hf : irreducible f) (hp : 0 n₁ = n₂ ∧ g₁ = g₂ := begin revert g₁ g₂, - wlog hn : n₁ ≤ n₂ := le_total n₁ n₂ using [n₁ n₂, n₂ n₁], + wlog hn : n₁ ≤ n₂, + { intros g₁ g₂ hg₁ Hg₁ hg₂ Hg₂, + simpa only [eq_comm] using this hf hp n₂ n₁ (le_of_not_le hn) g₂ g₁ hg₂ Hg₂ hg₁ Hg₁ }, have hf0 : f ≠ 0 := hf.ne_zero, unfreezingI { intros, rw le_iff_exists_add at hn, rcases hn with ⟨k, rfl⟩, rw [← hgf₁, pow_add, expand_mul, expand_inj (pow_pos hp n₁)] at hgf₂, subst hgf₂, @@ -344,8 +346,6 @@ begin { rw is_unit_iff at h, rcases h with ⟨r, hr, rfl⟩, simp_rw expand_C at hf, exact absurd (is_unit_C.2 hr) hf.1 }, { rw [add_zero, pow_zero, expand_one], split; refl } }, - obtain ⟨hn, hg⟩ := this g₂ g₁ hg₂ hgf₂ hg₁ hgf₁, - exact ⟨hn.symm, hg.symm⟩ end end char_p diff --git a/src/field_theory/separable_degree.lean b/src/field_theory/separable_degree.lean index e91513df84c43..31645af5673f8 100644 --- a/src/field_theory/separable_degree.lean +++ b/src/field_theory/separable_degree.lean @@ -109,7 +109,8 @@ theorem contraction_degree_eq_or_insep (hg : g.separable) (hg' : g'.separable) : g.nat_degree = g'.nat_degree := begin - wlog hm : m ≤ m' := le_total m m' using [m m' g g', m' m g' g], + wlog hm : m ≤ m', + { exact (this g' g m' m h_expand.symm hg' hg (le_of_not_le hm)).symm }, obtain ⟨s, rfl⟩ := exists_add_of_le hm, rw [pow_add, expand_mul, expand_inj (pow_pos (ne_zero.pos q) m)] at h_expand, subst h_expand, diff --git a/src/group_theory/order_of_element.lean b/src/group_theory/order_of_element.lean index f7b78b4c442b5..5c629adbebdea 100644 --- a/src/group_theory/order_of_element.lean +++ b/src/group_theory/order_of_element.lean @@ -410,7 +410,8 @@ by rw [modeq_zero_iff_dvd, order_of_dvd_iff_pow_eq_one] @[to_additive] lemma pow_eq_pow_iff_modeq : x ^ n = x ^ m ↔ n ≡ m [MOD (order_of x)] := begin - wlog hmn : m ≤ n, + wlog hmn : m ≤ n generalizing m n, + { rw [eq_comm, modeq.comm, this (le_of_not_le hmn)], }, obtain ⟨k, rfl⟩ := nat.exists_eq_add_of_le hmn, rw [← mul_one (x ^ m), pow_add, mul_left_cancel_iff, pow_eq_one_iff_modeq], exact ⟨λ h, nat.modeq.add_left _ h, λ h, nat.modeq.add_left_cancel' _ h⟩, diff --git a/src/group_theory/perm/cycle/basic.lean b/src/group_theory/perm/cycle/basic.lean index 8341e96b2727e..c5d400b19a151 100644 --- a/src/group_theory/perm/cycle/basic.lean +++ b/src/group_theory/perm/cycle/basic.lean @@ -588,7 +588,8 @@ begin obtain ⟨x, hx, -⟩ := id hf, exact ⟨x, hx, by simp [h]⟩ }, { rintro ⟨x, hx, hx'⟩, - wlog hab : a ≤ b, + wlog hab : a ≤ b generalizing a b, + { exact (this hx'.symm (le_of_not_le hab)).symm }, suffices : f ^ (b - a) = 1, { rw [pow_sub _ hab, mul_inv_eq_one] at this, rw this }, diff --git a/src/linear_algebra/exterior_algebra/basic.lean b/src/linear_algebra/exterior_algebra/basic.lean index c0d133a728194..fba425a184892 100644 --- a/src/linear_algebra/exterior_algebra/basic.lean +++ b/src/linear_algebra/exterior_algebra/basic.lean @@ -245,9 +245,11 @@ let F := (multilinear_map.mk_pi_algebra_fin R n (exterior_algebra R M)).comp_lin in { map_eq_zero_of_eq' := λ f x y hfxy hxy, begin rw [multilinear_map.comp_linear_map_apply, multilinear_map.mk_pi_algebra_fin_apply], - wlog h : x < y := lt_or_gt_of_ne hxy using x y, + clear F, + wlog h : x < y, + { exact this n f y x hfxy.symm hxy.symm (hxy.lt_or_lt.resolve_left h), }, clear hxy, - induction n with n hn generalizing x y, + induction n with n hn, { exact x.elim0, }, { rw [list.of_fn_succ, list.prod_cons], by_cases hx : x = 0, @@ -258,8 +260,8 @@ in -- ignore the left-most term and induct on the remaining ones, decrementing indices { convert mul_zero _, refine hn (λ i, f $ fin.succ i) - (x.pred hx) (y.pred (ne_of_lt $ lt_of_le_of_lt x.zero_le h).symm) - (fin.pred_lt_pred_iff.mpr h) _, + (x.pred hx) (y.pred (ne_of_lt $ lt_of_le_of_lt x.zero_le h).symm) _ + (fin.pred_lt_pred_iff.mpr h), simp only [fin.succ_pred], exact hfxy, } } end, diff --git a/src/measure_theory/constructions/borel_space.lean b/src/measure_theory/constructions/borel_space.lean index 295b8f72dacfc..73c77a38e0bf5 100644 --- a/src/measure_theory/constructions/borel_space.lean +++ b/src/measure_theory/constructions/borel_space.lean @@ -1363,16 +1363,13 @@ begin ennreal.coe_ne_top, preimage_Union, inter_Union], { assume i j, simp only [function.on_fun], - wlog h : i ≤ j := le_total i j using [i j, j i] tactic.skip, - { assume hij, - replace hij : i + 1 ≤ j := lt_of_le_of_ne h hij, - apply disjoint_left.2 (λ x hx h'x, lt_irrefl (f x) _), - calc f x < t ^ (i + 1) : hx.2.2 - ... ≤ t ^ j : ennreal.zpow_le_of_le (ennreal.one_le_coe_iff.2 ht.le) hij - ... ≤ f x : h'x.2.1 }, - { assume hij, - rw disjoint.comm, - exact this hij.symm } }, + assume hij, + wlog h : i < j generalizing i j, + { exact (this hij.symm (hij.lt_or_lt.resolve_left h)).symm }, + apply disjoint_left.2 (λ x hx h'x, lt_irrefl (f x) _), + calc f x < t ^ (i + 1) : hx.2.2 + ... ≤ t ^ j : ennreal.zpow_le_of_le (ennreal.one_le_coe_iff.2 ht.le) h + ... ≤ f x : h'x.2.1 }, { assume n, exact hs.inter (hf measurable_set_Ico) } }, rw [A, B, C, add_assoc], diff --git a/src/measure_theory/covering/besicovitch.lean b/src/measure_theory/covering/besicovitch.lean index c86b9f073a593..86423e2f90bc3 100644 --- a/src/measure_theory/covering/besicovitch.lean +++ b/src/measure_theory/covering/besicovitch.lean @@ -260,7 +260,8 @@ begin by_contra, suffices H : function.injective p.index, from not_injective_of_ordinal p.index H, assume x y hxy, - wlog x_le_y : x ≤ y := le_total x y using [x y, y x], + wlog x_le_y : x ≤ y generalizing x y, + { exact (this hxy.symm (le_of_not_le x_le_y)).symm }, rcases eq_or_lt_of_le x_le_y with rfl|H, { refl }, simp only [nonempty_def, not_exists, exists_prop, not_and, not_lt, not_le, mem_set_of_eq, not_forall] at h, @@ -395,16 +396,14 @@ begin rpos := λ k, p.rpos (p.index (G k)), h := begin assume a b a_ne_b, - wlog G_le : G a ≤ G b := le_total (G a) (G b) using [a b, b a] tactic.skip, - { have G_lt : G a < G b, - { rcases G_le.lt_or_eq with H|H, { exact H }, - have A : (a : ℕ) ≠ b := fin.coe_injective.ne a_ne_b, - rw [← color_G a (nat.lt_succ_iff.1 a.2), ← color_G b (nat.lt_succ_iff.1 b.2), H] at A, - exact (A rfl).elim }, - exact or.inl (Gab a b G_lt) }, - { assume a_ne_b, - rw or_comm, - exact this a_ne_b.symm } + wlog G_le : G a ≤ G b generalizing a b, + { exact (this b a a_ne_b.symm (le_of_not_le G_le)).symm }, + have G_lt : G a < G b, + { rcases G_le.lt_or_eq with H|H, { exact H }, + have A : (a : ℕ) ≠ b := fin.coe_injective.ne a_ne_b, + rw [← color_G a (nat.lt_succ_iff.1 a.2), ← color_G b (nat.lt_succ_iff.1 b.2), H] at A, + exact (A rfl).elim }, + exact or.inl (Gab a b G_lt), end, hlast := begin assume a ha, @@ -456,11 +455,8 @@ begin obtain ⟨jy, jy_lt, jyi, rfl⟩ : ∃ (jy : ordinal), jy < p.last_step ∧ p.color jy = i ∧ y = p.index jy, by simpa only [exists_prop, mem_Union, mem_singleton_iff] using hy, - wlog jxy : jx ≤ jy := le_total jx jy using [jx jy, jy jx] tactic.skip, - swap, - { assume h1 h2 h3 h4 h5 h6 h7, - rw [function.on_fun, disjoint.comm], - exact this h4 h5 h6 h1 h2 h3 h7.symm }, + wlog jxy : jx ≤ jy generalizing jx jy, + { exact (this jy jy_lt jyi hy jx jx_lt jxi hx x_ne_y.symm (le_of_not_le jxy)).symm }, replace jxy : jx < jy, by { rcases lt_or_eq_of_le jxy with H|rfl, { exact H }, { exact (x_ne_y rfl).elim } }, let A : set ℕ := ⋃ (j : {j // j < jy}) diff --git a/src/measure_theory/covering/besicovitch_vector_space.lean b/src/measure_theory/covering/besicovitch_vector_space.lean index dc64f7093c509..9cba2ebdeafdb 100644 --- a/src/measure_theory/covering/besicovitch_vector_space.lean +++ b/src/measure_theory/covering/besicovitch_vector_space.lean @@ -501,13 +501,11 @@ begin by_cases hi : ‖a.c i‖ = 0; field_simp [norm_smul, hi] }, refine ⟨c', λ n, norm_c'_le n, λ i j inej, _⟩, - -- up to exchanging `i` and `j`, one can assume `‖c i‖ ≤ ‖c j‖`. - wlog hij : ‖a.c i‖ ≤ ‖a.c j‖ := le_total (‖a.c i‖) (‖a.c j‖) using [i j, j i] tactic.skip, swap, - { assume i_ne_j, - rw norm_sub_rev, - exact this i_ne_j.symm }, + -- up to exchanging `i` and `j`, one can assume `∥c i∥ ≤ ∥c j∥`. + wlog hij : ‖a.c i‖ ≤ ‖a.c j‖ generalizing i j, + { rw norm_sub_rev, exact this j i inej.symm (le_of_not_le hij) }, rcases le_or_lt (‖a.c j‖) 2 with Hj|Hj, - -- case `‖c j‖ ≤ 2` (and therefore also `‖c i‖ ≤ 2`) + -- case `∥c j∥ ≤ 2` (and therefore also `∥c i∥ ≤ 2`) { simp_rw [c', Hj, hij.trans Hj, if_true], exact exists_normalized_aux1 a lastr hτ δ hδ1 hδ2 i j inej }, -- case `2 < ‖c j‖` diff --git a/src/measure_theory/integral/divergence_theorem.lean b/src/measure_theory/integral/divergence_theorem.lean index cb324a0182956..1e6b317345a3c 100644 --- a/src/measure_theory/integral/divergence_theorem.lean +++ b/src/measure_theory/integral/divergence_theorem.lean @@ -486,31 +486,31 @@ lemma integral2_divergence_prod_of_has_fderiv_within_at_off_countable (f g : ℝ (∫ x in a₁..b₁, g (x, b₂)) - (∫ x in a₁..b₁, g (x, a₂)) + (∫ y in a₂..b₂, f (b₁, y)) - ∫ y in a₂..b₂, f (a₁, y) := begin - wlog h₁ : a₁ ≤ b₁ := le_total a₁ b₁ using [a₁ b₁, b₁ a₁] tactic.skip, - wlog h₂ : a₂ ≤ b₂ := le_total a₂ b₂ using [a₂ b₂, b₂ a₂] tactic.skip, - { simp only [uIcc_of_le h₁, uIcc_of_le h₂, min_eq_left, max_eq_right, h₁, h₂] - at Hcf Hcg Hdf Hdg Hi, - calc ∫ x in a₁..b₁, ∫ y in a₂..b₂, f' (x, y) (1, 0) + g' (x, y) (0, 1) - = ∫ x in Icc a₁ b₁, ∫ y in Icc a₂ b₂, f' (x, y) (1, 0) + g' (x, y) (0, 1) : - by simp only [interval_integral.integral_of_le, h₁, h₂, - set_integral_congr_set_ae Ioc_ae_eq_Icc] - ... = ∫ x in Icc a₁ b₁ ×ˢ Icc a₂ b₂, f' x (1, 0) + g' x (0, 1) : - (set_integral_prod _ Hi).symm - ... = (∫ x in a₁..b₁, g (x, b₂)) - (∫ x in a₁..b₁, g (x, a₂)) + - (∫ y in a₂..b₂, f (b₁, y)) - ∫ y in a₂..b₂, f (a₁, y) : - begin - rw Icc_prod_Icc at *, - apply integral_divergence_prod_Icc_of_has_fderiv_within_at_off_countable_of_le f g f' g' - (a₁, a₂) (b₁, b₂) ⟨h₁, h₂⟩ s; assumption - end }, - { rw [uIcc_comm b₂ a₂, min_comm b₂ a₂, max_comm b₂ a₂] at this, - intros Hcf Hcg Hdf Hdg Hi, - simp only [interval_integral.integral_symm b₂ a₂, interval_integral.integral_neg], - refine (congr_arg has_neg.neg (this Hcf Hcg Hdf Hdg Hi)).trans _, abel }, - { rw [uIcc_comm b₁ a₁, min_comm b₁ a₁, max_comm b₁ a₁] at this, - intros Hcf Hcg Hdf Hdg Hi, + wlog h₁ : a₁ ≤ b₁ generalizing a₁ b₁, + { specialize this b₁ a₁, + rw [uIcc_comm b₁ a₁, min_comm b₁ a₁, max_comm b₁ a₁] at this, simp only [interval_integral.integral_symm b₁ a₁], - refine (congr_arg has_neg.neg (this Hcf Hcg Hdf Hdg Hi)).trans _, abel } + refine (congr_arg has_neg.neg (this Hcf Hcg Hdf Hdg Hi (le_of_not_le h₁))).trans _, abel }, + wlog h₂ : a₂ ≤ b₂ generalizing a₂ b₂, + { specialize this b₂ a₂, + rw [uIcc_comm b₂ a₂, min_comm b₂ a₂, max_comm b₂ a₂] at this, + simp only [interval_integral.integral_symm b₂ a₂, interval_integral.integral_neg], + refine (congr_arg has_neg.neg (this Hcf Hcg Hdf Hdg Hi (le_of_not_le h₂))).trans _, abel }, + simp only [uIcc_of_le h₁, uIcc_of_le h₂, min_eq_left, max_eq_right, h₁, h₂] + at Hcf Hcg Hdf Hdg Hi, + calc ∫ x in a₁..b₁, ∫ y in a₂..b₂, f' (x, y) (1, 0) + g' (x, y) (0, 1) + = ∫ x in Icc a₁ b₁, ∫ y in Icc a₂ b₂, f' (x, y) (1, 0) + g' (x, y) (0, 1) : + by simp only [interval_integral.integral_of_le, h₁, h₂, + set_integral_congr_set_ae Ioc_ae_eq_Icc] + ... = ∫ x in Icc a₁ b₁ ×ˢ Icc a₂ b₂, f' x (1, 0) + g' x (0, 1) : + (set_integral_prod _ Hi).symm + ... = (∫ x in a₁..b₁, g (x, b₂)) - (∫ x in a₁..b₁, g (x, a₂)) + + (∫ y in a₂..b₂, f (b₁, y)) - ∫ y in a₂..b₂, f (a₁, y) : + begin + rw Icc_prod_Icc at *, + apply integral_divergence_prod_Icc_of_has_fderiv_within_at_off_countable_of_le f g f' g' + (a₁, a₂) (b₁, b₂) ⟨h₁, h₂⟩ s; assumption + end end end measure_theory diff --git a/src/measure_theory/integral/interval_integral.lean b/src/measure_theory/integral/interval_integral.lean index 275b56a23c36d..95fd16215a161 100644 --- a/src/measure_theory/integral/interval_integral.lean +++ b/src/measure_theory/integral/interval_integral.lean @@ -437,8 +437,8 @@ by simpa only [mul_comm] using comp_mul_left hf c lemma comp_add_right (hf : interval_integrable f volume a b) (c : ℝ) : interval_integrable (λ x, f (x + c)) volume (a - c) (b - c) := begin - wlog h := le_total a b using [a b, b a] tactic.skip, - swap, { exact λ h, interval_integrable.symm (this h.symm) }, + wlog h : a ≤ b, + { exact interval_integrable.symm (this hf.symm _ (le_of_not_le h)) }, rw interval_integrable_iff' at hf ⊢, have A : measurable_embedding (λ x, x + c) := (homeomorph.add_right c).closed_embedding.measurable_embedding, @@ -982,12 +982,11 @@ by { rw [integral_interval_sub_interval_comm hab hcd hac, integral_symm b d, int lemma integral_Iic_sub_Iic (ha : integrable_on f (Iic a) μ) (hb : integrable_on f (Iic b) μ) : ∫ x in Iic b, f x ∂μ - ∫ x in Iic a, f x ∂μ = ∫ x in a..b, f x ∂μ := begin - wlog hab : a ≤ b using [a b] tactic.skip, - { rw [sub_eq_iff_eq_add', integral_of_le hab, ← integral_union (Iic_disjoint_Ioc le_rfl), - Iic_union_Ioc_eq_Iic hab], - exacts [measurable_set_Ioc, ha, hb.mono_set (λ _, and.right)] }, - { intros ha hb, - rw [integral_symm, ← this hb ha, neg_sub] } + wlog hab : a ≤ b generalizing a b, + { rw [integral_symm, ← this hb ha (le_of_not_le hab), neg_sub] }, + rw [sub_eq_iff_eq_add', integral_of_le hab, ← integral_union (Iic_disjoint_Ioc le_rfl), + Iic_union_Ioc_eq_Iic hab], + exacts [measurable_set_Ioc, ha, hb.mono_set (λ _, and.right)] end /-- If `μ` is a finite measure then `∫ x in a..b, c ∂μ = (μ (Iic b) - μ (Iic a)) • c`. -/ diff --git a/src/number_theory/number_field/embeddings.lean b/src/number_theory/number_field/embeddings.lean index 0abf10cf00fe2..55667d8484f5f 100644 --- a/src/number_theory/number_field/embeddings.lean +++ b/src/number_theory/number_field/embeddings.lean @@ -118,11 +118,9 @@ lemma pow_eq_one_of_norm_eq_one {x : K} begin obtain ⟨a, -, b, -, habne, h⟩ := @set.infinite.exists_ne_map_eq_of_maps_to _ _ _ _ ((^) x : ℕ → K) set.infinite_univ _ (finite_of_norm_le K A (1:ℝ)), - { replace habne := habne.lt_or_lt, - have : _, swap, cases habne, swap, - { revert a b, exact this }, - { exact this b a h.symm habne }, - refine λ a b h hlt, ⟨a - b, tsub_pos_of_lt hlt, _⟩, + { wlog hlt : b < a, + { exact this hxi hx b a habne.symm h.symm (habne.lt_or_lt.resolve_right hlt) }, + refine ⟨a - b, tsub_pos_of_lt hlt, _⟩, rw [← nat.sub_add_cancel hlt.le, pow_add, mul_left_eq_self₀] at h, refine h.resolve_right (λ hp, _), specialize hx (is_alg_closed.lift (number_field.is_algebraic K)).to_ring_hom, diff --git a/src/number_theory/padics/padic_norm.lean b/src/number_theory/padics/padic_norm.lean index 0380f329ff94a..4629482bd257e 100644 --- a/src/number_theory/padics/padic_norm.lean +++ b/src/number_theory/padics/padic_norm.lean @@ -192,7 +192,8 @@ and the norm of `q`. -/ protected theorem nonarchimedean {q r : ℚ} : padic_norm p (q + r) ≤ max (padic_norm p q) (padic_norm p r) := begin - wlog hle := le_total (padic_val_rat p q) (padic_val_rat p r) using [q r], + wlog hle : (padic_val_rat p q) ≤ (padic_val_rat p r) generalizing q r, + { rw [add_comm, max_comm], exact this (le_of_not_le hle) }, exact nonarchimedean_aux hle end @@ -213,8 +214,8 @@ of the norms of `q` and `r`. -/ lemma add_eq_max_of_ne {q r : ℚ} (hne : padic_norm p q ≠ padic_norm p r) : padic_norm p (q + r) = max (padic_norm p q) (padic_norm p r) := begin - wlog hle := le_total (padic_norm p r) (padic_norm p q) using [q r], - have hlt : padic_norm p r < padic_norm p q, from lt_of_le_of_ne hle hne.symm, + wlog hlt : (padic_norm p r) < (padic_norm p q), + { rw [add_comm, max_comm], exact this hne.symm (hne.lt_or_lt.resolve_right hlt) }, have : padic_norm p q ≤ max (padic_norm p (q + r)) (padic_norm p r), from calc padic_norm p q = padic_norm p (q + r - r) : by congr; ring ... ≤ max (padic_norm p (q + r)) (padic_norm p (-r)) : padic_norm.nonarchimedean diff --git a/src/order/omega_complete_partial_order.lean b/src/order/omega_complete_partial_order.lean index b75b68bae3ace..38b71e8c1dc8e 100644 --- a/src/order/omega_complete_partial_order.lean +++ b/src/order/omega_complete_partial_order.lean @@ -284,7 +284,7 @@ lemma eq_of_chain {c : chain (part α)} {a b : α} (ha : some a ∈ c) (hb : som begin cases ha with i ha, replace ha := ha.symm, cases hb with j hb, replace hb := hb.symm, - wlog h : i ≤ j := le_total i j using [a b i j, b a j i], + wlog h : i ≤ j, { exact (this j hb i ha (le_of_not_le h)).symm }, rw [eq_some_iff] at ha hb, have := c.monotone h _ ha, apply mem_unique this hb end diff --git a/src/ring_theory/roots_of_unity.lean b/src/ring_theory/roots_of_unity.lean index c21e28e689fb0..938d0c449973c 100644 --- a/src/ring_theory/roots_of_unity.lean +++ b/src/ring_theory/roots_of_unity.lean @@ -331,7 +331,8 @@ h.pow_ne_one_of_pos_of_lt zero_lt_one hk ∘ (pow_one ζ).trans lemma pow_inj (h : is_primitive_root ζ k) ⦃i j : ℕ⦄ (hi : i < k) (hj : j < k) (H : ζ ^ i = ζ ^ j) : i = j := begin - wlog hij : i ≤ j, + wlog hij : i ≤ j generalizing i j, + { exact (this hj hi H.symm (le_of_not_le hij)).symm }, apply le_antisymm hij, rw ← tsub_eq_zero_iff_le, apply nat.eq_zero_of_dvd_of_lt _ (lt_of_le_of_lt tsub_le_self hj), diff --git a/src/ring_theory/valuation/basic.lean b/src/ring_theory/valuation/basic.lean index e1dc335cac974..a32dd3133d0fc 100644 --- a/src/ring_theory/valuation/basic.lean +++ b/src/ring_theory/valuation/basic.lean @@ -238,15 +238,13 @@ begin suffices : ¬v (x + y) < max (v x) (v y), from or_iff_not_imp_right.1 (le_iff_eq_or_lt.1 (v.map_add x y)) this, intro h', - wlog vyx : v y < v x using x y, - { apply lt_or_gt_of_ne h.symm }, - { rw max_eq_left_of_lt vyx at h', - apply lt_irrefl (v x), - calc v x = v ((x+y) - y) : by simp - ... ≤ max (v $ x + y) (v y) : map_sub _ _ _ - ... < v x : max_lt h' vyx }, - { apply this h.symm, - rwa [add_comm, max_comm] at h' } + wlog vyx : v y < v x, + { refine this v h.symm _ (h.lt_or_lt.resolve_right vyx), rwa [add_comm, max_comm] }, + rw max_eq_left_of_lt vyx at h', + apply lt_irrefl (v x), + calc v x = v ((x+y) - y) : by simp + ... ≤ max (v $ x + y) (v y) : map_sub _ _ _ + ... < v x : max_lt h' vyx end lemma map_add_eq_of_lt_right (h : v x < v y) : v (x + y) = v y := diff --git a/src/set_theory/cardinal/divisibility.lean b/src/set_theory/cardinal/divisibility.lean index 768dc2e8895cd..c0c8df3ff4a99 100644 --- a/src/set_theory/cardinal/divisibility.lean +++ b/src/set_theory/cardinal/divisibility.lean @@ -69,6 +69,8 @@ begin cases eq_or_ne (b * c) 0 with hz hz, { rcases mul_eq_zero.mp hz with rfl | rfl; simp }, wlog h : c ≤ b, + { cases le_total c b; [skip, rw or_comm]; apply_assumption, assumption', + all_goals { rwa mul_comm } }, left, have habc := le_of_dvd hz hbc, rwa [mul_eq_max' $ ha.trans $ habc, max_def', if_pos h] at hbc @@ -113,8 +115,9 @@ begin { intro h, rw [h, zero_dvd_iff, mul_eq_zero] at hbc, cases hbc; contradiction }, - wlog hℵ₀ : ℵ₀ ≤ b := hℵ₀ using [b c], - exact or.inl (dvd_of_le_of_aleph_0_le hn ((nat_lt_aleph_0 n).le.trans hℵ₀) hℵ₀), + wlog hℵ₀b : ℵ₀ ≤ b, + { refine (this h c b _ _ hc hb hℵ₀.symm hn (hℵ₀.resolve_left hℵ₀b)).symm; rwa mul_comm }, + exact or.inl (dvd_of_le_of_aleph_0_le hn ((nat_lt_aleph_0 n).le.trans hℵ₀b) hℵ₀b), end lemma is_prime_iff {a : cardinal} : prime a ↔ ℵ₀ ≤ a ∨ ∃ p : ℕ, a = p ∧ p.prime := diff --git a/src/set_theory/game/nim.lean b/src/set_theory/game/nim.lean index d768c21db2dee..fbbbb91d3abbd 100644 --- a/src/set_theory/game/nim.lean +++ b/src/set_theory/game/nim.lean @@ -192,16 +192,13 @@ end @[simp] lemma nim_add_equiv_zero_iff (o₁ o₂ : ordinal) : nim o₁ + nim o₂ ≈ 0 ↔ o₁ = o₂ := begin split, - { refine not_imp_not.1 (λ (h : _ ≠ _), (impartial.not_equiv_zero_iff _).2 _), - obtain h | h := h.lt_or_lt, - { rw [impartial.fuzzy_zero_iff_gf, zero_lf_le, nim_def o₂], - refine ⟨to_left_moves_add (sum.inr _), _⟩, - { exact (ordinal.principal_seg_out h).top }, - { simpa using (impartial.add_self (nim o₁)).2 } }, - { rw [impartial.fuzzy_zero_iff_gf, zero_lf_le, nim_def o₁], - refine ⟨to_left_moves_add (sum.inl _), _⟩, - { exact (ordinal.principal_seg_out h).top }, - { simpa using (impartial.add_self (nim o₂)).2 } } }, + { refine not_imp_not.1 (λ (hne : _ ≠ _), (impartial.not_equiv_zero_iff _).2 _), + wlog h : o₁ < o₂, + { exact (fuzzy_congr_left add_comm_equiv).1 (this _ _ hne.symm (hne.lt_or_lt.resolve_left h)) }, + rw [impartial.fuzzy_zero_iff_gf, zero_lf_le, nim_def o₂], + refine ⟨to_left_moves_add (sum.inr _), _⟩, + { exact (ordinal.principal_seg_out h).top }, + { simpa using (impartial.add_self (nim o₁)).2 } }, { rintro rfl, exact impartial.add_self (nim o₁) } end diff --git a/src/set_theory/surreal/dyadic.lean b/src/set_theory/surreal/dyadic.lean index fa55c3c1c7a2d..e663ebaecd75e 100644 --- a/src/set_theory/surreal/dyadic.lean +++ b/src/set_theory/surreal/dyadic.lean @@ -169,6 +169,7 @@ lemma dyadic_aux {m₁ m₂ : ℤ} {y₁ y₂ : ℕ} (h₂ : m₁ * (2 ^ y₁) = begin revert m₁ m₂, wlog h : y₁ ≤ y₂, + { intros m₁ m₂ aux, exact (this (le_of_not_le h) aux.symm).symm }, intros m₁ m₂ h₂, obtain ⟨c, rfl⟩ := le_iff_exists_add.mp h, rw [add_comm, pow_add, ← mul_assoc, mul_eq_mul_right_iff] at h₂, diff --git a/src/tactic/wlog.lean b/src/tactic/wlog.lean index 67af4ce1411fc..618096020358c 100644 --- a/src/tactic/wlog.lean +++ b/src/tactic/wlog.lean @@ -1,242 +1,85 @@ /- Copyright (c) 2018 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Johannes Hölzl - -Without loss of generality tactic. +Authors: Johannes Hölzl, Mario Carneiro, Johan Commelin, Reid Barton -/ -import data.list.perm - -open expr -setup_tactic_parser -namespace tactic +import tactic.core +import tactic.dependencies -private meta def update_pp_name : expr → name → expr -| (local_const n _ bi d) pp := local_const n pp bi d -| e n := e +/-! -private meta def elim_or : ℕ → expr → tactic (list expr) -| 0 h := fail "zero cases" -| 1 h := return [h] -| (n + 1) h := do - [(_, [hl], []), (_, [hr], [])] ← induction h, -- there should be no dependent terms - [gl, gr] ← get_goals, - set_goals [gr], - hsr ← elim_or n hr, - gsr ← get_goals, - set_goals (gl :: gsr), - return (hl :: hsr) +# Without loss of generality tactic -private meta def dest_or : expr → tactic (list expr) | e := do - `(%%a ∨ %%b) ← whnf e | return [e], - lb ← dest_or b, - return (a :: lb) - -private meta def match_perms (pat : pattern) : expr → tactic (list $ list expr) | t := - (do - m ← match_pattern pat t, - guard (m.2.all expr.is_local_constant), - return [m.2]) <|> - (do - `(%%l ∨ %%r) ← whnf t, - m ← match_pattern pat l, - rs ← match_perms r, - return (m.2 :: rs)) - -meta def wlog (vars' : list expr) (h_cases fst_case : expr) (perms : list (list expr)) : - tactic unit := do - guard h_cases.is_local_constant, +The tactic `wlog h : P` will add an assumption `h : P` to the main goal, +and add a new goal that requires showing that the case `h : ¬ P` can be reduced to the case +where `P` holds (typically by symmetry). +The new goal will be placed at the top of the goal stack. - -- reorder s.t. context is Γ ⬝ vars ⬝ cases ⊢ ∀deps, … - nr ← revert_lst (vars' ++ [h_cases]), - vars ← intron' vars'.length, - h_cases ← intro h_cases.local_pp_name, - - cases ← infer_type h_cases, - h_fst_case ← - mk_local_def h_cases.local_pp_name - (fst_case.instantiate_locals $ (vars'.zip vars).map $ λ⟨o, n⟩, (o.local_uniq_name, n)), - ((), pr) ← solve_aux cases (repeat $ exact h_fst_case <|> left >> skip), - - t ← target, - fixed_vars ← vars.mmap update_type, - let t' := (instantiate_local h_cases.local_uniq_name pr t).pis (fixed_vars ++ [h_fst_case]), - - (h, [g]) ← local_proof `this t' (do - clear h_cases, - vars.mmap clear, - intron nr), - - h₀ :: hs ← elim_or perms.length h_cases, +-/ - solve1 (do - exact (h.mk_app $ vars ++ [h₀])), +namespace tactic - focus ((hs.zip perms.tail).map $ λ⟨h_case, perm⟩, do - let p_v := (vars'.zip vars).map (λ⟨p, v⟩, (p.local_uniq_name, v)), - let p := perm.map (λp, p.instantiate_locals p_v), - note `this none (h.mk_app $ p ++ [h_case]), - clear h, - return ()), - gs ← get_goals, - set_goals (g :: gs) +/-- A helper function to retrieve the names of the first `n` arguments to a Pi-expression. -/ +meta def take_pi_args : nat → expr → list name +| (n+1) (expr.pi h _ _ e) := h :: take_pi_args n e +| _ _ := [] namespace interactive -open interactive interactive.types expr - -private meta def parse_permutations : option (list (list name)) → tactic (list (list expr)) -| none := return [] -| (some []) := return [] -| (some perms@(p₀ :: ps)) := do - (guard p₀.nodup <|> fail - "No permutation `xs_i` in `using [xs_1, …, xs_n]` should contain the same variable twice."), - (guard (perms.all $ λp, p.perm p₀) <|> - fail ("The permutations `xs_i` in `using [xs_1, …, xs_n]` must be permutations of the same" ++ - " variables.")), - perms.mmap (λp, p.mmap get_local) - -/-- Without loss of generality: reduces to one goal under variables permutations. - -Given a goal of the form `g xs`, a predicate `p` over a set of variables, as well as variable -permutations `xs_i`. Then `wlog` produces goals of the form - -* The case goal, i.e. the permutation `xs_i` covers all possible cases: - `⊢ p xs_0 ∨ ⋯ ∨ p xs_n` -* The main goal, i.e. the goal reduced to `xs_0`: - `(h : p xs_0) ⊢ g xs_0` -* The invariant goals, i.e. `g` is invariant under `xs_i`: - `(h : p xs_i) (this : g xs_0) ⊢ gs xs_i` - -Either the permutation is provided, or a proof of the disjunction is provided to compute the -permutation. The disjunction need to be in assoc normal form, e.g. `p₀ ∨ (p₁ ∨ p₂)`. In many cases -the invariant goals can be solved by AC rewriting using `cc` etc. - -For example, on a state `(n m : ℕ) ⊢ p n m` the tactic `wlog h : n ≤ m using [n m, m n]` produces -the following states: -* `(n m : ℕ) ⊢ n ≤ m ∨ m ≤ n` -* `(n m : ℕ) (h : n ≤ m) ⊢ p n m` -* `(n m : ℕ) (h : m ≤ n) (this : p n m) ⊢ p m n` - -`wlog` supports different calling conventions. The name `h` is used to give a name to the introduced -case hypothesis. If the name is avoided, the default will be `case`. - -1. `wlog : p xs0 using [xs0, …, xsn]` - Results in the case goal `p xs0 ∨ ⋯ ∨ ps xsn`, the main goal `(case : p xs0) ⊢ g xs0` and the - invariance goals `(case : p xsi) (this : g xs0) ⊢ g xsi`. - -2. `wlog : p xs0 := r using xs0` - The expression `r` is a proof of the shape `p xs0 ∨ ⋯ ∨ p xsi`, it is also used to compute the - variable permutations. - -3. `wlog := r using xs0` - The expression `r` is a proof of the shape `p xs0 ∨ ⋯ ∨ p xsi`, it is also used to compute the - variable permutations. This is not as stable as (2), for example `p` cannot be a disjunction. +setup_tactic_parser -4. `wlog : R x y using x y` and `wlog : R x y` - Produces the case `R x y ∨ R y x`. If `R` is ≤, then the disjunction discharged using linearity. - If `using x y` is avoided then `x` and `y` are the last two variables appearing in the - expression `R x y`. --/ -meta def wlog - (h : parse ident?) - (pat : parse (tk ":" *> texpr)?) - (cases : parse (tk ":=" *> texpr)?) - (perms : parse (tk "using" *> (list_of (ident*) <|> (λx, [x]) <$> ident*))?) - (discharger : tactic unit := - (tactic.solve_by_elim <|> tactic.tautology {classical := tt} <|> - using_smt (smt_tactic.intros >> smt_tactic.solve_goals))) : +/-- `wlog h : P` will add an assumption `h : P` to the main goal, +and add a side goal that requires showing that the case `h : ¬ P` can be reduced to the case +where `P` holds (typically by symmetry). + +The side goal will be at the top of the stack. In this side goal, there will be two assumptions: +- `h : ¬ P`: the assumption that `P` does not hold +- `this`: which is the statement that in the old context `P` suffices to prove the goal. + By default, the name `this` is used, but the idiom `with H` can be added to specify the name: + `wlog h : P with H`. + +Typically, it is useful to use the variant `wlog h : P generalizing x y`, +to revert certain parts of the context before creating the new goal. +In this way, the wlog-claim `this` can be applied to `x` and `y` in different orders +(exploiting symmetry, which is the typical use case). + +By default, the entire context is reverted. -/ +meta def wlog (H : parse ident) (t : parse (tk ":" *> texpr)) + (revert : parse ((tk "generalizing" *> ((none <$ tk "*") <|> some <$> ident*)) <|> pure none)) + (h : parse (tk "with" *> ident)?) : tactic unit := do -perms ← parse_permutations perms, -(pat, cases_pr, cases_goal, vars, perms) ← (match cases with -| some r := do - vars::_ ← return perms | - fail "At least one set of variables expected, i.e. `using x y` or `using [x y, y x]`.", - cases_pr ← to_expr r, - cases_pr ← (if cases_pr.is_local_constant - then return $ match h with some n := update_pp_name cases_pr n | none := cases_pr end - else do - note (h.get_or_else `case) none cases_pr), - cases ← infer_type cases_pr, - (pat, perms') ← match pat with - | some pat := do - pat ← to_expr pat, - let vars' := vars.filter $ λv, v.occurs pat, - case_pat ← tactic.mk_pattern [] vars' pat [] vars', - perms' ← match_perms case_pat cases, - return (pat, perms') - | none := do - (p :: ps) ← dest_or cases, - let vars' := vars.filter $ λv, v.occurs p, - case_pat ← tactic.mk_pattern [] vars' p [] vars', - perms' ← (p :: ps).mmap (λp, do m ← match_pattern case_pat p, return m.2), - return (p, perms') - end, - let vars_name := vars.map local_uniq_name, - guard (perms'.all $ λp, p.all $ λv, v.is_local_constant ∧ v.local_uniq_name ∈ vars_name) <|> - fail "Cases contains variables not declared in `using x y z`", - perms ← (if perms.length = 1 - then do - return (perms'.map $ λ p, - p ++ vars.filter (λ v, p.all (λ v', v'.local_uniq_name ≠ v.local_uniq_name))) - else do - guard (perms.length = perms'.length) <|> - fail "The provided permutation list has a different length then the provided cases.", - return perms), - return (pat, cases_pr, @none expr, vars, perms) - -| none := do - let name_h := h.get_or_else `case, - some pat ← return pat | fail "Either specify cases or a pattern with permutations", - pat ← to_expr pat, - (do - [x, y] ← match perms with - | [] := return pat.list_local_consts - | [l] := return l - | _ := failed + -- if there is no `with` clause, use `this` as default name + let h := h.get_or_else `this, + t ← i_to_expr ``(%%t : Sort*), + -- compute which constants must be reverted (by default: everything) + (num_generalized, goal, rctx) ← retrieve (do + assert_core H t, swap, + -- use `revert_lst'` to ensure that the order of local constants in the context is preserved + num_generalized ← match revert with + | none := revert_all + | some revert := prod.fst <$> (revert.mmap tactic.get_local >>= revert_lst') end, - let cases := mk_or_lst - [pat, pat.instantiate_locals [(x.local_uniq_name, y), (y.local_uniq_name, x)]], - (do - `(%%x' ≤ %%y') ← return pat, - (cases_pr, []) ← local_proof name_h cases (exact ``(le_total %%x' %%y')), - return (pat, cases_pr, none, [x, y], [[x, y], [y, x]])) - <|> - (do - (cases_pr, [g]) ← local_proof name_h cases skip, - return (pat, cases_pr, some g, [x, y], [[x, y], [y, x]]))) <|> - (do - guard (perms.length ≥ 2) <|> - fail ("To generate cases at least two permutations are required, i.e. `using [x y, y x]`" ++ - " or exactly 0 or 2 variables"), - (vars :: perms') ← return perms, - let names := vars.map local_uniq_name, - let cases := mk_or_lst (pat :: perms'.map (λp, pat.instantiate_locals (names.zip p))), - (cases_pr, [g]) ← local_proof name_h cases skip, - return (pat, cases_pr, some g, vars, perms)) -end), -let name_fn := if perms.length = 2 then λ _, `invariant else - λ i, mk_simple_name ("invariant_" ++ to_string (i + 1)), -with_enable_tags $ tactic.focus1 $ do - t ← get_main_tag, - tactic.wlog vars cases_pr pat perms, - tactic.focus (set_main_tag (mk_num_name `_case 0 :: `main :: t) :: - (list.range (perms.length - 1)).map (λi, do - set_main_tag (mk_num_name `_case 0 :: name_fn i :: t), - try discharger)), - match cases_goal with - | some g := do - set_tag g (mk_num_name `_case 0 :: `cases :: t), - gs ← get_goals, - set_goals (g :: gs) - | none := skip - end + goal ← target, + ctx ← local_context, + return (num_generalized, goal, ctx)), + ctx ← local_context, + e ← tactic.assert h goal, + goal ← target, + (take_pi_args num_generalized goal).reverse.mmap' $ λ h, + try (tactic.get_local h >>= tactic.clear), + intron (num_generalized + 1), + -- prove the easy branch of the side goal + swap, + tactic.by_cases t H, + H ← tactic.get_local H, + let L := ctx.filter (λ n, n ∉ rctx), + tactic.exact $ (e.mk_app L).app H add_tactic_doc -{ name := "wlog", - category := doc_category.tactic, - decl_names := [``wlog], - tags := ["logic"] } +{ name := "wlog", + category := doc_category.tactic, + decl_names := [`tactic.interactive.wlog], + tags := ["logic"] } end interactive diff --git a/src/topology/algebra/order/left_right_lim.lean b/src/topology/algebra/order/left_right_lim.lean index be928f32253b8..6628268721fd9 100644 --- a/src/topology/algebra/order/left_right_lim.lean +++ b/src/topology/algebra/order/left_right_lim.lean @@ -253,15 +253,13 @@ begin have fs_count : (f '' s).countable, { have A : (f '' s).pairwise_disjoint (λ x, Ioo x (z (inv_fun_on f s x))), { rintros _ ⟨u, us, rfl⟩ _ ⟨v, vs, rfl⟩ huv, - wlog h'uv : u ≤ v := le_total u v using [u v, v u] tactic.skip, - { rcases eq_or_lt_of_le h'uv with rfl|h''uv, - { exact (huv rfl).elim }, - apply disjoint_iff_forall_ne.2, - rintros a ha b hb rfl, - simp [I.left_inv_on_inv_fun_on us, I.left_inv_on_inv_fun_on vs] at ha hb, - exact lt_irrefl _ ((ha.2.trans_le ((hz u us).2 v h''uv)).trans hb.1) }, - { assume hu hv h'uv, - exact (this hv hu h'uv.symm).symm } }, + wlog hle : u ≤ v generalizing u v, + { exact (this v vs u us huv.symm (le_of_not_le hle)).symm }, + have hlt : u < v, from hle.lt_of_ne (ne_of_apply_ne _ huv), + apply disjoint_iff_forall_ne.2, + rintros a ha b hb rfl, + simp only [I.left_inv_on_inv_fun_on us, I.left_inv_on_inv_fun_on vs] at ha hb, + exact lt_irrefl _ ((ha.2.trans_le ((hz u us).2 v hlt)).trans hb.1) }, apply set.pairwise_disjoint.countable_of_Ioo A, rintros _ ⟨y, ys, rfl⟩, simpa only [I.left_inv_on_inv_fun_on ys] using (hz y ys).1 }, diff --git a/src/topology/algebra/with_zero_topology.lean b/src/topology/algebra/with_zero_topology.lean index 16f04ede602c6..a0c786b819171 100644 --- a/src/topology/algebra/with_zero_topology.lean +++ b/src/topology/algebra/with_zero_topology.lean @@ -163,9 +163,9 @@ instance : has_continuous_mul Γ₀ := ⟨begin rw continuous_iff_continuous_at, rintros ⟨x, y⟩, - wlog hle : x ≤ y := le_total x y using [x y, y x] tactic.skip, swap, - { simpa only [mul_comm, (∘), prod.swap] - using tendsto.comp this (continuous_swap.tendsto (x, y)) }, + wlog hle : x ≤ y generalizing x y, + { have := tendsto.comp (this y x (le_of_not_le hle)) (continuous_swap.tendsto (x,y)), + simpa only [mul_comm, function.comp, prod.swap], }, rcases eq_or_ne x 0 with rfl|hx; [rcases eq_or_ne y 0 with rfl|hy, skip], { rw [continuous_at, zero_mul], refine ((has_basis_nhds_zero.prod_nhds has_basis_nhds_zero).tendsto_iff has_basis_nhds_zero).2 diff --git a/src/topology/connected.lean b/src/topology/connected.lean index 0dab6c95e6b2f..b261cc203a9f2 100644 --- a/src/topology/connected.lean +++ b/src/topology/connected.lean @@ -94,7 +94,9 @@ theorem is_preconnected_of_forall {s : set α} (x : α) begin rintros u v hu hv hs ⟨z, zs, zu⟩ ⟨y, ys, yv⟩, have xs : x ∈ s, by { rcases H y ys with ⟨t, ts, xt, yt, ht⟩, exact ts xt }, - wlog xu : x ∈ u := hs xs using [u v y z, v u z y], + wlog xu : x ∈ u, + { rw inter_comm u v, rw union_comm at hs, + exact this x H v u hv hu hs y ys yv z zs zu xs ((hs xs).resolve_right xu), }, rcases H y ys with ⟨t, ts, xt, yt, ht⟩, have := ht u v hu hv(subset.trans ts hs) ⟨x, xt, xu⟩ ⟨y, yt, yv⟩, exact this.imp (λ z hz, ⟨ts hz.1, hz.2⟩) diff --git a/src/topology/instances/rat_lemmas.lean b/src/topology/instances/rat_lemmas.lean index 641939a78f8ba..fe0059dec2da3 100644 --- a/src/topology/instances/rat_lemmas.lean +++ b/src/topology/instances/rat_lemmas.lean @@ -79,7 +79,8 @@ instance : totally_disconnected_space ℚ := begin refine ⟨λ s hsu hs x hx y hy, _⟩, clear hsu, by_contra' H : x ≠ y, - wlog hlt : x < y := H.lt_or_lt using [x y, y x], + wlog hlt : x < y, + { exact this s hs y hy x hx H.symm (H.lt_or_lt.resolve_left hlt) }, rcases exists_irrational_btwn (rat.cast_lt.2 hlt) with ⟨z, hz, hxz, hzy⟩, have := hs.image coe continuous_coe_real.continuous_on, rw [is_preconnected_iff_ord_connected] at this, diff --git a/src/topology/metric_space/basic.lean b/src/topology/metric_space/basic.lean index c1364354d899c..b74e0400d871c 100644 --- a/src/topology/metric_space/basic.lean +++ b/src/topology/metric_space/basic.lean @@ -1527,12 +1527,10 @@ 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, `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] } + wlog h : b ≤ a, + { rw [nndist_comm, max_comm], exact this b a (le_of_not_le h) }, + 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], end @[simp] lemma nnreal.nndist_zero_eq_val (z : ℝ≥0) : nndist 0 z = z := diff --git a/src/topology/metric_space/emetric_paracompact.lean b/src/topology/metric_space/emetric_paracompact.lean index e12583452acaf..554e959682fde 100644 --- a/src/topology/metric_space/emetric_paracompact.lean +++ b/src/topology/metric_space/emetric_paracompact.lean @@ -135,8 +135,9 @@ begin -- For each `m ≤ n + k` there is at most one `j` such that `D m j ∩ B` is nonempty. have Hle : ∀ m ≤ n + k, set.subsingleton {j | (D m j ∩ B).nonempty}, { rintros m hm j₁ ⟨y, hyD, hyB⟩ j₂ ⟨z, hzD, hzB⟩, - by_contra h, - wlog h : j₁ < j₂ := ne.lt_or_lt h using [j₁ j₂ y z, j₂ j₁ z y], + by_contra' h' : j₁ ≠ j₂, + wlog h : j₁ < j₂ generalizing j₁ j₂ y z, + { exact this z hzD hzB y hyD hyB h'.symm (h'.lt_or_lt.resolve_left h), }, rcases memD.1 hyD with ⟨y', rfl, hsuby, -, hdisty⟩, rcases memD.1 hzD with ⟨z', rfl, -, -, hdistz⟩, suffices : edist z' y' < 3 * 2⁻¹ ^ m, from nmem_of_lt_ind h (hsuby this), diff --git a/src/topology/separation.lean b/src/topology/separation.lean index a4ef2b889c4ec..ac51baa7da286 100644 --- a/src/topology/separation.lean +++ b/src/topology/separation.lean @@ -208,7 +208,9 @@ theorem minimal_nonempty_closed_subsingleton [t0_space α] {s : set α} (hs : is begin refine λ x hx y hy, of_not_not (λ hxy, _), rcases exists_is_open_xor_mem hxy with ⟨U, hUo, hU⟩, - wlog h : x ∈ U ∧ y ∉ U := hU using [x y, y x], cases h with hxU hyU, + wlog h : x ∈ U ∧ y ∉ U, + { exact this hmin y hy x hx (ne.symm hxy) U hUo hU.symm (hU.resolve_left h), }, + cases h with hxU hyU, have : s \ U = s := hmin (s \ U) (diff_subset _ _) ⟨y, hy, hyU⟩ (hs.sdiff hUo), exact (this.symm.subset hx).2 hxU end @@ -235,7 +237,9 @@ theorem minimal_nonempty_open_subsingleton [t0_space α] {s : set α} (hs : is_o begin refine λ x hx y hy, of_not_not (λ hxy, _), rcases exists_is_open_xor_mem hxy with ⟨U, hUo, hU⟩, - wlog h : x ∈ U ∧ y ∉ U := hU using [x y, y x], cases h with hxU hyU, + wlog h : x ∈ U ∧ y ∉ U, + { exact this hs hmin y hy x hx (ne.symm hxy) U hUo hU.symm (hU.resolve_left h), }, + cases h with hxU hyU, have : s ∩ U = s := hmin (s ∩ U) (inter_subset_left _ _) ⟨x, hx, hxU⟩ (hs.inter hUo), exact hyU (this.symm.subset hy).2 end @@ -1474,11 +1478,11 @@ instance t3_space.t2_5_space [t3_space α] : t2_5_space α := begin refine ⟨λ x y hne, _⟩, rw [lift'_nhds_closure, lift'_nhds_closure], - have : x ∉ closure {y} ∨ y ∉ closure {x}, + have aux : x ∉ closure {y} ∨ y ∉ closure {x}, from (t0_space_iff_or_not_mem_closure α).mp infer_instance x y hne, - wlog H : x ∉ closure {y} := this using [x y, y x] tactic.skip, + wlog H : x ∉ closure ({y} : set α), + { refine (this y x hne.symm aux.symm (aux.resolve_left H)).symm }, { rwa [← disjoint_nhds_nhds_set, nhds_set_singleton] at H }, - { exact λ h, (this h.symm).symm } end protected lemma embedding.t3_space [topological_space β] [t3_space β] {f : α → β} diff --git a/src/topology/shrinking_lemma.lean b/src/topology/shrinking_lemma.lean index 9224bb9a6526a..94ca1524a6844 100644 --- a/src/topology/shrinking_lemma.lean +++ b/src/topology/shrinking_lemma.lean @@ -94,7 +94,8 @@ lemma apply_eq_of_chain {c : set (partial_refinement u s)} (hc : is_chain (≤) (h₁ : v₁ ∈ c) (h₂ : v₂ ∈ c) {i} (hi₁ : i ∈ v₁.carrier) (hi₂ : i ∈ v₂.carrier) : v₁ i = v₂ i := begin - wlog hle : v₁ ≤ v₂ := hc.total h₁ h₂ using [v₁ v₂, v₂ v₁], + wlog hle : v₁ ≤ v₂, + { cases hc.total h₁ h₂; [skip, symmetry]; apply_assumption; assumption' }, exact hle.2 _ hi₁, end diff --git a/test/wlog.lean b/test/wlog.lean index 71a3917b6bcbf..f93d507f48a51 100644 --- a/test/wlog.lean +++ b/test/wlog.lean @@ -1,131 +1,46 @@ /- Copyright (c) 2018 Simon Hudon. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Simon Hudon +Authors: Simon Hudon, Johan Commelin -/ import tactic.wlog section wlog -example {x y : ℕ} (a : x = 1) : true := -begin - suffices : false, trivial, - wlog h : x = y, - { guard_target x = y ∨ y = x, - admit }, - { guard_hyp h : x = y, - guard_hyp a : x = 1, - admit } -end - example {x y : ℕ} : true := begin - suffices : false, trivial, wlog h : x ≤ y, + { guard_hyp h : ¬x ≤ y, + guard_hyp this : ∀ {x y : ℕ}, x ≤ y → true, -- `wlog` generalizes by default + guard_target true, + trivial }, { guard_hyp h : x ≤ y, - guard_target false, - admit } -end - -example {x y z : ℕ} : true := -begin - suffices : false, trivial, - wlog : x ≤ y + z using x y, - { guard_target x ≤ y + z ∨ y ≤ x + z, - admit }, - { guard_hyp case : x ≤ y + z, - guard_target false, - admit }, -end - -example {x : ℕ} (S₀ S₁ : set ℕ) (P : ℕ → Prop) - (h : x ∈ S₀ ∪ S₁) : true := -begin - suffices : false, trivial, - wlog h' : x ∈ S₀ using S₀ S₁, - { guard_target x ∈ S₀ ∨ x ∈ S₁, - admit }, - { guard_hyp h : x ∈ S₀ ∪ S₁, - guard_hyp h' : x ∈ S₀, - admit } -end - -example {n m i : ℕ} {p : ℕ → ℕ → ℕ → Prop} : true := -begin - suffices : false, trivial, - wlog : p n m i using [n m i, n i m, i n m], - { guard_target p n m i ∨ p n i m ∨ p i n m, - admit }, - { guard_hyp case : p n m i, - admit } + guard_target true, + trivial }, end -example {n m i : ℕ} {p : ℕ → Prop} : true := -begin - suffices : false, trivial, - wlog : p n using [n m i, m n i, i n m], - { guard_target p n ∨ p m ∨ p i, - admit }, - { guard_hyp case : p n, - admit } -end - -example {n m i : ℕ} {p : ℕ → ℕ → Prop} {q : ℕ → ℕ → ℕ → Prop} : true := -begin - suffices : q n m i, trivial, - have h : p n i ∨ p i m ∨ p m i, from sorry, - wlog : p n i := h using n m i, - { guard_hyp h : p n i, - guard_target q n m i, - admit }, - { guard_hyp h : p i m, - guard_hyp this : q i m n, - guard_target q n m i, - admit }, - { guard_hyp h : p m i, - guard_hyp this : q m i n, - guard_target q n m i, - admit }, -end - -example (X : Type) (A B C : set X) : A ∩ (B ∪ C) = (A ∩ B) ∪ (A ∩ C) := -begin - ext x, - split, - { intro hyp, - cases hyp, - wlog x_in : x ∈ B using B C, - { assumption }, - { exact or.inl ⟨hyp_left, x_in⟩ } }, - { intro hyp, - wlog x_in : x ∈ A ∩ B using B C, - { assumption }, - { exact ⟨x_in.left, or.inl x_in.right⟩ } } -end - -example (X : Type) (A B C : set X) : A ∩ (B ∪ C) = (A ∩ B) ∪ (A ∩ C) := +example {x y : ℕ} : true := begin - ext x, - split, - { intro hyp, - wlog x_in : x ∈ B := hyp.2 using B C, - { exact or.inl ⟨hyp.1, x_in⟩ } }, - { intro hyp, - wlog x_in : x ∈ A ∩ B := hyp using B C, - { exact ⟨x_in.left, or.inl x_in.right⟩ } } + wlog h : x ≤ y generalizing x, + { guard_hyp h : ¬x ≤ y, + guard_hyp this : ∀ {x : ℕ}, x ≤ y → true, -- only `x` was generalized + guard_target true, + trivial }, + { guard_hyp h : x ≤ y, + guard_target true, + trivial }, end -example (X : Type) (A B C : set X) : A ∩ (B ∪ C) = (A ∩ B) ∪ (A ∩ C) := +example {x y z : ℕ} : true := begin - ext x, - split, - { intro hyp, - cases hyp, - wlog x_in : x ∈ B := hyp_right using B C, - { exact or.inl ⟨hyp_left, x_in⟩ }, }, - { intro hyp, - wlog x_in : x ∈ A ∩ B := hyp using B C, - { exact ⟨x_in.left, or.inl x_in.right⟩ } } + wlog h : x ≤ y + z with H, + { guard_hyp h : ¬ x ≤ y + z, + guard_hyp H : ∀ {x y z : ℕ}, x ≤ y + z → true, -- wlog-claim is named `H` instead of `this` + guard_target true, + trivial }, + { guard_hyp h : x ≤ y + z, + guard_target true, + trivial }, end end wlog From 00ed28cdf9d7de148a4a1ff1158e56773fe9d6bc Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Fri, 3 Feb 2023 13:16:57 +0000 Subject: [PATCH 0420/1291] chore(100.yaml): Brouwer fix point (#18372) @Shamrock-Frost formalized Brouwer's fix point theorem a few months ago, and we never got around to adding this to the list of 100 theorems. This should be the 75th entry if I counted correctly :tada: --- docs/100.yaml | 3 +++ 1 file changed, 3 insertions(+) diff --git a/docs/100.yaml b/docs/100.yaml index 53980901664be..408d2f1487e0c 100644 --- a/docs/100.yaml +++ b/docs/100.yaml @@ -133,6 +133,9 @@ author : Moritz Doll 36: title : Brouwer Fixed Point Theorem + author : Brendan Seamas Murphy + links : + result: https://github.com/Shamrock-Frost/BrouwerFixedPoint/blob/master/src/brouwer_fixed_point.lean 37: title : The Solution of a Cubic author : Jeoff Lee From b7399344324326918d65d0c74e9571e3a8cb9199 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 3 Feb 2023 18:30:21 +0000 Subject: [PATCH 0421/1291] feat(algebra/order/chebyshev): Special case using division (#18369) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit along with its application to the proof of Szemerédi Regularity Lemma. --- src/algebra/order/chebyshev.lean | 13 +++++++ .../simple_graph/regularity/bound.lean | 36 +++++++++++++++++++ 2 files changed, 49 insertions(+) diff --git a/src/algebra/order/chebyshev.lean b/src/algebra/order/chebyshev.lean index c825cf8a4f8b4..6a1bdef1eae39 100644 --- a/src/algebra/order/chebyshev.lean +++ b/src/algebra/order/chebyshev.lean @@ -128,3 +128,16 @@ lemma antivary.card_mul_sum_le_sum_mul_sum (hfg : antivary f g) : (hfg.antivary_on _).card_mul_sum_le_sum_mul_sum end mul + +variables [linear_ordered_field α] {s : finset ι} {f : ι → α} + +lemma sum_div_card_sq_le_sum_sq_div_card : + ((∑ i in s, f i) / s.card) ^ 2 ≤ (∑ i in s, f i ^ 2) / s.card := +begin + obtain rfl | hs := s.eq_empty_or_nonempty, + { simp }, + rw [←card_pos, ←@nat.cast_pos α] at hs, + rw [div_pow, div_le_div_iff (sq_pos_of_ne_zero _ hs.ne') hs, sq (s.card : α), mul_left_comm, + ←mul_assoc], + exact mul_le_mul_of_nonneg_right (sq_sum_le_card_mul_sum_sq) hs.le, +end diff --git a/src/combinatorics/simple_graph/regularity/bound.lean b/src/combinatorics/simple_graph/regularity/bound.lean index b90bf7744f082..285654851ff71 100644 --- a/src/combinatorics/simple_graph/regularity/bound.lean +++ b/src/combinatorics/simple_graph/regularity/bound.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 Yaël Dillies, Bhavik Mehta. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies, Bhavik Mehta -/ +import algebra.order.chebyshev import analysis.special_functions.pow import order.partition.equipartition @@ -11,6 +12,8 @@ import order.partition.equipartition This file gathers the numerical facts required by the proof of Szemerédi's regularity lemma. +This entire file is internal to the proof of Szemerédi Regularity Lemma. + ## Main declarations * `szemeredi_regularity.step_bound`: During the inductive step, a partition of size `n` is blown to @@ -21,6 +24,7 @@ This file gathers the numerical facts required by the proof of Szemerédi's regu -/ open finset fintype function real +open_locale big_operators namespace szemeredi_regularity @@ -143,6 +147,38 @@ lemma initial_bound_le_bound : initial_bound ε l ≤ bound ε l := lemma le_bound : l ≤ bound ε l := (le_initial_bound ε l).trans $ initial_bound_le_bound ε l lemma bound_pos : 0 < bound ε l := (initial_bound_pos ε l).trans_le $ initial_bound_le_bound ε l +variables {ι 𝕜 : Type*} [linear_ordered_field 𝕜] (r : ι → ι → Prop) [decidable_rel r] + {s t : finset ι} {x : 𝕜} + +lemma mul_sq_le_sum_sq (hst : s ⊆ t) (f : ι → 𝕜) (hs : x^2 ≤ ((∑ i in s, f i) / s.card) ^ 2) + (hs' : (s.card : 𝕜) ≠ 0) : + (s.card : 𝕜) * x ^ 2 ≤ ∑ i in t, f i ^ 2 := +(mul_le_mul_of_nonneg_left (hs.trans sum_div_card_sq_le_sum_sq_div_card) $ + nat.cast_nonneg _).trans $ (mul_div_cancel' _ hs').le.trans $ sum_le_sum_of_subset_of_nonneg hst $ + λ i _ _, sq_nonneg _ + +lemma add_div_le_sum_sq_div_card (hst : s ⊆ t) (f : ι → 𝕜) (d : 𝕜) (hx : 0 ≤ x) + (hs : x ≤ |(∑ i in s, f i)/s.card - (∑ i in t, f i)/t.card|) + (ht : d ≤ ((∑ i in t, f i)/t.card)^2) : + d + s.card/t.card * x^2 ≤ (∑ i in t, f i^2)/t.card := +begin + obtain hscard | hscard := (s.card.cast_nonneg : (0 : 𝕜) ≤ s.card).eq_or_lt, + { simpa [←hscard] using ht.trans sum_div_card_sq_le_sum_sq_div_card }, + have htcard : (0:𝕜) < t.card := hscard.trans_le (nat.cast_le.2 (card_le_of_subset hst)), + have h₁ : x^2 ≤ ((∑ i in s, f i)/s.card - (∑ i in t, f i)/t.card)^2 := + sq_le_sq.2 (by rwa [abs_of_nonneg hx]), + have h₂ : x^2 ≤ ((∑ i in s, (f i - (∑ j in t, f j)/t.card))/s.card)^2, + { apply h₁.trans, + rw [sum_sub_distrib, sum_const, nsmul_eq_mul, sub_div, mul_div_cancel_left _ hscard.ne'] }, + apply (add_le_add_right ht _).trans, + rw [←mul_div_right_comm, le_div_iff htcard, add_mul, div_mul_cancel _ htcard.ne'], + have h₃ := mul_sq_le_sum_sq hst (λ i, f i - (∑ j in t, f j) / t.card) h₂ hscard.ne', + apply (add_le_add_left h₃ _).trans, + simp [←mul_div_right_comm _ (t.card : 𝕜), sub_div' _ _ _ htcard.ne', ←sum_div, ←add_div, mul_pow, + div_le_iff (sq_pos_of_ne_zero _ htcard.ne'), sub_sq, sum_add_distrib, ←sum_mul, ←mul_sum], + ring_nf, +end + end szemeredi_regularity namespace tactic From a4ec43f53b0bd44c697bcc3f5a62edd56f269ef1 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 3 Feb 2023 18:30:22 +0000 Subject: [PATCH 0422/1291] refactor(combinatorics/simple_graph/density): Generalize to arbitrary ordered fields (#18370) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit I originally thought making graph densities `ℚ`-valued would work best, but I now need it to be a real and it turns out to be more pain than anything. --- src/combinatorics/simple_graph/density.lean | 32 ++++++++++++--------- 1 file changed, 18 insertions(+), 14 deletions(-) diff --git a/src/combinatorics/simple_graph/density.lean b/src/combinatorics/simple_graph/density.lean index c57129e7e3f8b..df0129197b6ce 100644 --- a/src/combinatorics/simple_graph/density.lean +++ b/src/combinatorics/simple_graph/density.lean @@ -24,14 +24,14 @@ Between two finsets of vertices, open finset open_locale big_operators -variables {ι κ α β : Type*} +variables {𝕜 ι κ α β : Type*} /-! ### Density of a relation -/ namespace rel section asymmetric -variables (r : α → β → Prop) [Π a, decidable_pred (r a)] {s s₁ s₂ : finset α} {t t₁ t₂ : finset β} - {a : α} {b : β} {δ : ℚ} +variables [linear_ordered_field 𝕜] (r : α → β → Prop) [Π a, decidable_pred (r a)] + {s s₁ s₂ : finset α} {t t₁ t₂ : finset β} {a : α} {b : β} {δ : 𝕜} /-- Finset of edges of a relation between two finsets of vertices. -/ def interedges (s : finset α) (t : finset β) : finset (α × β) := (s ×ˢ t).filter $ λ e, r e.1 e.2 @@ -188,32 +188,36 @@ end lemma abs_edge_density_sub_edge_density_le_two_mul_sub_sq (hs : s₂ ⊆ s₁) (ht : t₂ ⊆ t₁) (hδ₀ : 0 ≤ δ) (hδ₁ : δ < 1) (hs₂ : (1 - δ) * s₁.card ≤ s₂.card) (ht₂ : (1 - δ) * t₁.card ≤ t₂.card) : - |edge_density r s₂ t₂ - edge_density r s₁ t₁| ≤ 2*δ - δ^2 := + |(edge_density r s₂ t₂ : 𝕜) - edge_density r s₁ t₁| ≤ 2*δ - δ^2 := begin have hδ' : 0 ≤ 2 * δ - δ ^ 2, { rw [sub_nonneg, sq], exact mul_le_mul_of_nonneg_right (hδ₁.le.trans (by norm_num)) hδ₀ }, rw ←sub_pos at hδ₁, - simp only [edge_density], obtain rfl | hs₂' := s₂.eq_empty_or_nonempty, { rw [finset.card_empty, nat.cast_zero] at hs₂, - simpa [(nonpos_of_mul_nonpos_right hs₂ hδ₁).antisymm (nat.cast_nonneg _)] using hδ' }, + simpa [edge_density, (nonpos_of_mul_nonpos_right hs₂ hδ₁).antisymm (nat.cast_nonneg _)] + using hδ' }, obtain rfl | ht₂' := t₂.eq_empty_or_nonempty, { rw [finset.card_empty, nat.cast_zero] at ht₂, - simpa [(nonpos_of_mul_nonpos_right ht₂ hδ₁).antisymm (nat.cast_nonneg _)] using hδ' }, + simpa [edge_density, (nonpos_of_mul_nonpos_right ht₂ hδ₁).antisymm (nat.cast_nonneg _)] + using hδ' }, rw [show 2 * δ - δ ^ 2 = 1 - (1 - δ) * (1 - δ), by ring], - refine (abs_edge_density_sub_edge_density_le_one_sub_mul r hs ht hs₂' ht₂').trans _, - apply sub_le_sub_left (mul_le_mul ((le_div_iff _).2 hs₂) ((le_div_iff _).2 ht₂) hδ₁.le _), - { exact_mod_cast (hs₂'.mono hs).card_pos }, - { exact_mod_cast (ht₂'.mono ht).card_pos }, - { positivity } + norm_cast, + refine (rat.cast_le.2 $ + abs_edge_density_sub_edge_density_le_one_sub_mul r hs ht hs₂' ht₂').trans _, + push_cast, + have := hs₂'.mono hs, + have := ht₂'.mono ht, + refine sub_le_sub_left (mul_le_mul ((le_div_iff _).2 hs₂) ((le_div_iff _).2 ht₂) hδ₁.le _) _; + positivity, end /-- If `s₂ ⊆ s₁`, `t₂ ⊆ t₁` and they take up all but a `δ`-proportion, then the difference in edge densities is at most `2 * δ`. -/ lemma abs_edge_density_sub_edge_density_le_two_mul (hs : s₂ ⊆ s₁) (ht : t₂ ⊆ t₁) (hδ : 0 ≤ δ) (hscard : (1 - δ) * s₁.card ≤ s₂.card) (htcard : (1 - δ) * t₁.card ≤ t₂.card) : - |edge_density r s₂ t₂ - edge_density r s₁ t₁| ≤ 2 * δ := + |(edge_density r s₂ t₂ : 𝕜) - edge_density r s₁ t₁| ≤ 2 * δ := begin cases lt_or_le δ 1, { exact (abs_edge_density_sub_edge_density_le_two_mul_sub_sq r hs ht hδ h hscard htcard).trans @@ -324,7 +328,7 @@ lemma edge_density_add_edge_density_compl (hs : s.nonempty) (ht : t.nonempty) (h begin rw [edge_density_def, edge_density_def, div_add_div_same, div_eq_one_iff_eq], { exact_mod_cast card_interedges_add_card_interedges_compl _ h }, - { exact_mod_cast (mul_pos hs.card_pos ht.card_pos).ne' } + { positivity } end end decidable_eq From 85e3c05a94b27c84dc6f234cf88326d5e0096ec3 Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Fri, 3 Feb 2023 19:48:02 +0000 Subject: [PATCH 0423/1291] feat(ring_theory/ideal/norm): relative ideal norm (#18303) This PR contains the definition of `ideal.rel_norm`, the relative ideal norm as a bundled monoid-with-zero homomorphism sending `I : ideal S` to the ideal in `R` spanned by the norms of the elements in `I`; here `R` and `S` are Dedekind domains and `S` is an extension of `R` which is finite free as an `R`-module. Co-Authored-By: Alex J. Best Co-authored-by: Vierkantor --- src/algebra/char_p/quotient.lean | 12 +++ src/ring_theory/ideal/norm.lean | 169 +++++++++++++++++++++++++++++-- 2 files changed, 172 insertions(+), 9 deletions(-) diff --git a/src/algebra/char_p/quotient.lean b/src/algebra/char_p/quotient.lean index 12b4c9e2c689b..922b91b1b5740 100644 --- a/src/algebra/char_p/quotient.lean +++ b/src/algebra/char_p/quotient.lean @@ -37,3 +37,15 @@ lemma quotient' {R : Type*} [comm_ring R] (p : ℕ) [char_p R p] (I : ideal R) end⟩ end char_p + +lemma ideal.quotient.index_eq_zero {R : Type*} [comm_ring R] (I : ideal R) : + (I.to_add_subgroup.index : R ⧸ I) = 0 := +begin + rw [add_subgroup.index, nat.card_eq], + split_ifs with hq, swap, simp, + by_contra h, + -- TODO: can we avoid rewriting the `I.to_add_subgroup` here? + letI : fintype (R ⧸ I) := @fintype.of_finite _ hq, + have h : (fintype.card (R ⧸ I) : R ⧸ I) ≠ 0 := h, + simpa using h +end diff --git a/src/ring_theory/ideal/norm.lean b/src/ring_theory/ideal/norm.lean index 9b7e4ad6aa35b..158190d30f825 100644 --- a/src/ring_theory/ideal/norm.lean +++ b/src/ring_theory/ideal/norm.lean @@ -4,16 +4,15 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Anne Baanen, Alex J. Best -/ +import algebra.char_p.quotient import data.finsupp.fintype import data.int.absolute_value import data.int.associated -import data.matrix.notation -import data.zmod.quotient +import number_theory.ramification_inertia import linear_algebra.free_module.determinant import linear_algebra.free_module.ideal_quotient -import linear_algebra.free_module.pid -import linear_algebra.isomorphisms -import ring_theory.dedekind_domain.ideal +import ring_theory.dedekind_domain.pid +import ring_theory.local_properties import ring_theory.localization.norm /-! @@ -31,8 +30,10 @@ the norms of elements in `I`. This maps `⊥` to `0` and `⊤` to `1`. * `ideal.abs_norm (I : ideal R)`: the absolute ideal norm, defined as the cardinality of the quotient `R ⧸ I`, as a bundled monoid-with-zero homomorphism. - * `ideal.span_norm R (I : ideal S)`: the relative ideal norm, defined as - the ideal spanned by the norms of elements in `I`. + * `ideal.span_norm R (I : ideal S)`: the ideal spanned by the norms of elements in `I`. + This is used to define `ideal.rel_norm`. + * `ideal.rel_norm R (I : ideal S)`: the relative ideal norm as a bundled monoid-with-zero morphism, + defined as the ideal spanned by the norms of elements in `I`. ## Main results @@ -42,6 +43,7 @@ the norms of elements in `I`. of the basis change matrix * `ideal.abs_norm_span_singleton`: the ideal norm of a principal ideal is the norm of its generator + * `map_mul ideal.rel_norm`: multiplicativity of the relative ideal norm -/ open_locale big_operators @@ -82,6 +84,8 @@ end end submodule +section ring_of_integers + variables {S : Type*} [comm_ring S] [is_domain S] open submodule @@ -383,6 +387,12 @@ end lemma abs_norm_dvd_abs_norm_of_le {I J : ideal S} (h : J ≤ I) : I.abs_norm ∣ J.abs_norm := map_dvd abs_norm (dvd_iff_le.mpr h) +lemma abs_norm_dvd_norm_of_mem {I : ideal S} {x : S} (h : x ∈ I) : ↑I.abs_norm ∣ algebra.norm ℤ x := +begin + rw [← int.dvd_nat_abs, ← abs_norm_span_singleton x, int.coe_nat_dvd], + exact abs_norm_dvd_abs_norm_of_le ((span_singleton_le_iff_mem _).mpr h) +end + @[simp] lemma abs_norm_span_insert (r : S) (s : set S) : abs_norm (span (insert r s)) ∣ gcd (abs_norm (span s)) (algebra.norm ℤ r).nat_abs := @@ -410,8 +420,14 @@ lemma prime_of_irreducible_abs_norm_span {a : S} (ha : a ≠ 0) prime a := (ideal.span_singleton_prime ha).mp (is_prime_of_irreducible_abs_norm hI) +lemma abs_norm_mem (I : ideal S) : ↑I.abs_norm ∈ I := +by rw [abs_norm_apply, card_quot, ← ideal.quotient.eq_zero_iff_mem, map_nat_cast, + quotient.index_eq_zero] + end ideal +end ring_of_integers + end abs_norm section span_norm @@ -423,6 +439,8 @@ open submodule variables (R : Type*) [comm_ring R] {S : Type*} [comm_ring S] [algebra R S] /-- `ideal.span_norm R (I : ideal S)` is the ideal generated by mapping `algebra.norm R` over `I`. + +See also `ideal.rel_norm`. -/ def span_norm (I : ideal S) : ideal R := ideal.span (algebra.norm R '' (I : set S)) @@ -432,6 +450,21 @@ ideal.span (algebra.norm R '' (I : set S)) span_norm R (⊥ : ideal S) = ⊥ := span_eq_bot.mpr (λ x hx, by simpa using hx) +variables {R} + +@[simp] lemma span_norm_eq_bot_iff [is_domain R] [is_domain S] + [module.free R S] [module.finite R S] {I : ideal S} : + span_norm R I = ⊥ ↔ I = ⊥ := +begin + simp only [span_norm, ideal.span_eq_bot, set.mem_image, set_like.mem_coe, forall_exists_index, + and_imp, forall_apply_eq_imp_iff₂, + algebra.norm_eq_zero_iff_of_basis (module.free.choose_basis R S), @eq_bot_iff _ _ _ I, + set_like.le_def], + refl +end + +variables (R) + lemma norm_mem_span_norm {I : ideal S} (x : S) (hx : x ∈ I) : algebra.norm R x ∈ I.span_norm R := subset_span (set.mem_image_of_mem _ hx) @@ -444,7 +477,7 @@ le_antisymm end)) ((span_singleton_le_iff_mem _).mpr (norm_mem_span_norm _ _ (mem_span_singleton_self _))) -@[simp] lemma span_norm_top : span_norm R (⊤ : ideal S) = 1 := +@[simp] lemma span_norm_top : span_norm R (⊤ : ideal S) = ⊤ := by simp [← ideal.span_singleton_one] lemma map_span_norm (I : ideal S) {T : Type*} [comm_ring T] (f : R →+* T) : @@ -456,7 +489,7 @@ lemma span_norm_mono {I J : ideal S} (h : I ≤ J) : span_norm R I ≤ span_norm ideal.span_mono (set.monotone_image h) lemma span_norm_localization (I : ideal S) [module.finite R S] [module.free R S] - (M : submonoid R) {Rₘ Sₘ : Type*} + (M : submonoid R) {Rₘ : Type*} (Sₘ : Type*) [comm_ring Rₘ] [algebra R Rₘ] [comm_ring Sₘ] [algebra S Sₘ] [algebra Rₘ Sₘ] [algebra R Sₘ] [is_scalar_tower R Rₘ Sₘ] [is_scalar_tower R S Sₘ] [is_localization M Rₘ] [is_localization (algebra.algebra_map_submonoid S M) Sₘ] : @@ -490,6 +523,124 @@ begin all_goals { apply_instance } }, end +lemma span_norm_mul_span_norm_le (I J : ideal S) : + span_norm R I * span_norm R J ≤ span_norm R (I * J) := +begin + rw [span_norm, span_norm, span_norm, ideal.span_mul_span', ← set.image_mul], + refine ideal.span_mono (set.monotone_image _), + rintros _ ⟨x, y, hxI, hyJ, rfl⟩, + exact ideal.mul_mem_mul hxI hyJ +end + +/-- This condition `eq_bot_or_top` is equivalent to being a field. +However, `span_norm_mul_of_field` is harder to apply since we'd need to upgrade a `comm_ring R` +instance to a `field R` instance. -/ +lemma span_norm_mul_of_bot_or_top [is_domain R] [is_domain S] + [module.free R S] [module.finite R S] + (eq_bot_or_top : ∀ I : ideal R, I = ⊥ ∨ I = ⊤) + (I J : ideal S) : + span_norm R (I * J) = span_norm R I * span_norm R J := +begin + refine le_antisymm _ (span_norm_mul_span_norm_le _ _ _), + cases eq_bot_or_top (span_norm R I) with hI hI, + { rw [hI, span_norm_eq_bot_iff.mp hI, bot_mul, span_norm_bot], + exact bot_le }, + rw [hI, ideal.top_mul], + cases eq_bot_or_top (span_norm R J) with hJ hJ, + { rw [hJ, span_norm_eq_bot_iff.mp hJ, mul_bot, span_norm_bot], + exact bot_le }, + rw hJ, + exact le_top +end + +@[simp] lemma span_norm_mul_of_field {K : Type*} [field K] [algebra K S] [is_domain S] + [module.finite K S] (I J : ideal S) : + span_norm K (I * J) = span_norm K I * span_norm K J := +span_norm_mul_of_bot_or_top K eq_bot_or_top I J + +variables [is_domain R] [is_domain S] [is_dedekind_domain R] [is_dedekind_domain S] +variables [module.finite R S] [module.free R S] + +/-- Multiplicativity of `ideal.span_norm`. simp-normal form is `map_mul (ideal.rel_norm R)`. -/ +lemma span_norm_mul (I J : ideal S) : span_norm R (I * J) = span_norm R I * span_norm R J := +begin + nontriviality R, + casesI subsingleton_or_nontrivial S, + { have : ∀ I : ideal S, I = ⊤ := λ I, subsingleton.elim I ⊤, + simp [this I, this J, this (I * J)] }, + refine eq_of_localization_maximal _, + unfreezingI { intros P hP }, + by_cases hP0 : P = ⊥, + { unfreezingI { subst hP0 }, + rw span_norm_mul_of_bot_or_top, + intros I, + refine or_iff_not_imp_right.mpr (λ hI, _), + exact (hP.eq_of_le hI bot_le).symm }, + let P' := algebra.algebra_map_submonoid S P.prime_compl, + letI : algebra (localization.at_prime P) (localization P') := + localization_algebra P.prime_compl S, + haveI : is_scalar_tower R (localization.at_prime P) (localization P') := + is_scalar_tower.of_algebra_map_eq (λ x, (is_localization.map_eq _ _).symm), + have h : P' ≤ S⁰ := + map_le_non_zero_divisors_of_injective _ (no_zero_smul_divisors.algebra_map_injective _ _) + P.prime_compl_le_non_zero_divisors, + haveI : is_domain (localization P') := is_localization.is_domain_localization h, + haveI : is_dedekind_domain (localization P') := is_localization.is_dedekind_domain S h _, + letI := classical.dec_eq (ideal (localization P')), + haveI : is_principal_ideal_ring (localization P') := + is_dedekind_domain.is_principal_ideal_ring_localization_over_prime S P hP0, + rw [ideal.map_mul, ← span_norm_localization R I P.prime_compl (localization P'), + ← span_norm_localization R J P.prime_compl (localization P'), + ← span_norm_localization R (I * J) P.prime_compl (localization P'), ideal.map_mul, + ← (I.map _).span_singleton_generator, ← (J.map _).span_singleton_generator, + span_singleton_mul_span_singleton, span_norm_singleton, span_norm_singleton, + span_norm_singleton, span_singleton_mul_span_singleton, _root_.map_mul], + repeat { apply_instance }, + repeat { assumption }, +end + +/-- The relative norm `ideal.rel_norm R (I : ideal S)`, where `R` and `S` are Dedekind domains, +and `S` is an extension of `R` that is finite and free as a module. -/ +def rel_norm : ideal S →*₀ ideal R := +{ to_fun := span_norm R, + map_zero' := span_norm_bot R, + map_one' := by rw [one_eq_top, span_norm_top R, one_eq_top], + map_mul' := span_norm_mul R } + +lemma rel_norm_apply (I : ideal S) : + rel_norm R I = span (algebra.norm R '' (I : set S) : set R) := +rfl + +@[simp] lemma span_norm_eq (I : ideal S) : span_norm R I = rel_norm R I := rfl + +@[simp] lemma rel_norm_bot : rel_norm R (⊥ : ideal S) = ⊥ := +by simpa only [zero_eq_bot] using map_zero (rel_norm R : ideal S →*₀ _) + +@[simp] lemma rel_norm_top : rel_norm R (⊤ : ideal S) = ⊤ := +by simpa only [one_eq_top] using map_one (rel_norm R : ideal S →*₀ _) + +variables {R} + +@[simp] lemma rel_norm_eq_bot_iff {I : ideal S} : rel_norm R I = ⊥ ↔ I = ⊥ := +span_norm_eq_bot_iff + +variables (R) + +lemma norm_mem_rel_norm (I : ideal S) {x : S} (hx : x ∈ I) : algebra.norm R x ∈ rel_norm R I := +norm_mem_span_norm R x hx + +@[simp] lemma rel_norm_singleton (r : S) : + rel_norm R (span ({r} : set S)) = span {algebra.norm R r} := +span_norm_singleton R + +lemma map_rel_norm (I : ideal S) {T : Type*} [comm_ring T] (f : R →+* T) : + map f (rel_norm R I) = span ((f ∘ algebra.norm R) '' (I : set S)) := +map_span_norm R I f + +@[mono] +lemma rel_norm_mono {I J : ideal S} (h : I ≤ J) : rel_norm R I ≤ rel_norm R J := +span_norm_mono R h + end ideal end span_norm From b363547b3113d350d053abdf2884e9850a56b205 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 3 Feb 2023 19:48:04 +0000 Subject: [PATCH 0424/1291] chore(topology/algebra/infinite_sum): add lemmas about division (#18351) This also flips the direction of some misnamed statements. --- src/number_theory/l_series.lean | 2 +- src/topology/algebra/infinite_sum.lean | 29 +++++++++++++++++--------- 2 files changed, 20 insertions(+), 11 deletions(-) diff --git a/src/number_theory/l_series.lean b/src/number_theory/l_series.lean index c40f8bde46176..006b39f8105e5 100644 --- a/src/number_theory/l_series.lean +++ b/src/number_theory/l_series.lean @@ -58,7 +58,7 @@ begin simp [hf] }, refine summable_of_norm_bounded (λ (n : ℕ), m / (n ^ z)) _ _, { simp_rw [div_eq_mul_inv], - exact (summable_mul_left_iff h0).1 (real.summable_nat_rpow_inv.2 hz) }, + exact (summable_mul_left_iff h0).2 (real.summable_nat_rpow_inv.2 hz) }, { intro n, have hm : 0 ≤ m := le_trans (complex.abs.nonneg _) (h 0), cases n, diff --git a/src/topology/algebra/infinite_sum.lean b/src/topology/algebra/infinite_sum.lean index 653f696bcd0f1..069f4e8498ecd 100644 --- a/src/topology/algebra/infinite_sum.lean +++ b/src/topology/algebra/infinite_sum.lean @@ -1120,29 +1120,38 @@ by simp only [div_eq_mul_inv, h.mul_right b⁻¹] lemma summable.div_const (h : summable f) (b : α) : summable (λ x, f x / b) := (h.has_sum.div_const b).summable -lemma has_sum_mul_left_iff (h : a₂ ≠ 0) : has_sum f a₁ ↔ has_sum (λb, a₂ * f b) (a₂ * a₁) := -⟨has_sum.mul_left _, λ H, by simpa only [inv_mul_cancel_left₀ h] using H.mul_left a₂⁻¹⟩ +lemma has_sum_mul_left_iff (h : a₂ ≠ 0) : has_sum (λb, a₂ * f b) (a₂ * a₁) ↔ has_sum f a₁ := +⟨λ H, by simpa only [inv_mul_cancel_left₀ h] using H.mul_left a₂⁻¹, has_sum.mul_left _⟩ -lemma has_sum_mul_right_iff (h : a₂ ≠ 0) : has_sum f a₁ ↔ has_sum (λb, f b * a₂) (a₁ * a₂) := -⟨has_sum.mul_right _, λ H, by simpa only [mul_inv_cancel_right₀ h] using H.mul_right a₂⁻¹⟩ +lemma has_sum_mul_right_iff (h : a₂ ≠ 0) : has_sum (λb, f b * a₂) (a₁ * a₂) ↔ has_sum f a₁ := +⟨λ H, by simpa only [mul_inv_cancel_right₀ h] using H.mul_right a₂⁻¹, has_sum.mul_right _⟩ -lemma summable_mul_left_iff (h : a ≠ 0) : summable f ↔ summable (λb, a * f b) := -⟨λ H, H.mul_left _, λ H, by simpa only [inv_mul_cancel_left₀ h] using H.mul_left a⁻¹⟩ +lemma has_sum_div_const_iff (h : a₂ ≠ 0) : has_sum (λb, f b / a₂) (a₁ / a₂) ↔ has_sum f a₁ := +by simpa only [div_eq_mul_inv] using has_sum_mul_right_iff (inv_ne_zero h) -lemma summable_mul_right_iff (h : a ≠ 0) : summable f ↔ summable (λb, f b * a) := -⟨λ H, H.mul_right _, λ H, by simpa only [mul_inv_cancel_right₀ h] using H.mul_right a⁻¹⟩ +lemma summable_mul_left_iff (h : a ≠ 0) : summable (λb, a * f b) ↔ summable f := +⟨λ H, by simpa only [inv_mul_cancel_left₀ h] using H.mul_left a⁻¹, λ H, H.mul_left _⟩ + +lemma summable_mul_right_iff (h : a ≠ 0) : summable (λb, f b * a) ↔ summable f := +⟨λ H, by simpa only [mul_inv_cancel_right₀ h] using H.mul_right a⁻¹, λ H, H.mul_right _⟩ + +lemma summable_div_const_iff (h : a ≠ 0) : summable (λb, f b / a) ↔ summable f := +by simpa only [div_eq_mul_inv] using summable_mul_right_iff (inv_ne_zero h) lemma tsum_mul_left [t2_space α] : (∑' x, a * f x) = a * ∑' x, f x := if hf : summable f then hf.tsum_mul_left a else if ha : a = 0 then by simp [ha] else by rw [tsum_eq_zero_of_not_summable hf, - tsum_eq_zero_of_not_summable (mt (summable_mul_left_iff ha).2 hf), mul_zero] + tsum_eq_zero_of_not_summable (mt (summable_mul_left_iff ha).mp hf), mul_zero] lemma tsum_mul_right [t2_space α] : (∑' x, f x * a) = (∑' x, f x) * a := if hf : summable f then hf.tsum_mul_right a else if ha : a = 0 then by simp [ha] else by rw [tsum_eq_zero_of_not_summable hf, - tsum_eq_zero_of_not_summable (mt (summable_mul_right_iff ha).2 hf), zero_mul] + tsum_eq_zero_of_not_summable (mt (summable_mul_right_iff ha).mp hf), zero_mul] + +lemma tsum_div_const [t2_space α] : (∑' x, f x / a) = (∑' x, f x) / a := +by simpa only [div_eq_mul_inv] using tsum_mul_right end division_ring From ff6bde6f523f6228d5d3d90ffd915624da72a2ab Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Fri, 3 Feb 2023 21:38:21 +0000 Subject: [PATCH 0425/1291] feat(number_theory/number_field/embeddings): add card_complex_embeddings (#17753) Define complex and real places of a number field $K$ and prove that the number of real places is $r_1$ and the number of complex places is $r_2$ where $(r_1, r_2)$ is the signature of $K$. --- .../number_field/embeddings.lean | 115 +++++++++++++++++- 1 file changed, 113 insertions(+), 2 deletions(-) diff --git a/src/number_theory/number_field/embeddings.lean b/src/number_theory/number_field/embeddings.lean index 55667d8484f5f..e5446b71d7756 100644 --- a/src/number_theory/number_field/embeddings.lean +++ b/src/number_theory/number_field/embeddings.lean @@ -5,9 +5,9 @@ Authors: Alex J. Best, Xavier Roblot -/ import analysis.complex.polynomial -import data.complex.basic import field_theory.minpoly.is_integrally_closed import number_theory.number_field.basic +import ring_theory.norm import topology.instances.complex /-! @@ -321,12 +321,123 @@ begin { rw ← complex_embedding.is_real_conjugate_iff at hφ, rwa ← h, }}, { exact λ h, ⟨embedding w, h, mk_embedding w⟩, }, - end +end lemma not_is_real_iff_is_complex {w : infinite_place K} : ¬ is_real w ↔ is_complex w := by rw [is_complex_iff, is_real_iff] +/-- For `w` a real infinite place, return the corresponding embedding as a morphism `K →+* ℝ`. -/ +noncomputable def is_real.embedding {w : infinite_place K} (hw : is_real w) : K →+* ℝ := +(is_real_iff.mp hw).embedding + +@[simp] +lemma is_real.place_embedding_apply {w : infinite_place K} (hw : is_real w) (x : K): + place (is_real.embedding hw) x = w x := +begin + rw [is_real.embedding, complex_embedding.is_real.place_embedding, ← coe_mk], + exact congr_fun (congr_arg coe_fn (mk_embedding w)) x, +end + +variable (K) + +/-- The map from real embeddings to real infinite places as an equiv -/ +noncomputable def mk_real : + {φ : K →+* ℂ // complex_embedding.is_real φ} ≃ {w : infinite_place K // is_real w} := +{ to_fun := subtype.map mk (λ φ hφ, ⟨φ, hφ, rfl⟩), + inv_fun := λ w, ⟨w.1.embedding, is_real_iff.1 w.2⟩, + left_inv := λ φ, subtype.ext_iff.2 (number_field.complex_embeddings.is_real.embedding_mk φ.2), + right_inv := λ w, subtype.ext_iff.2 (mk_embedding w.1), } + +/-- The map from nonreal embeddings to complex infinite places -/ +noncomputable def mk_complex : + {φ : K →+* ℂ // ¬ complex_embedding.is_real φ} → {w : infinite_place K // is_complex w} := +subtype.map mk (λ φ hφ, ⟨φ, hφ, rfl⟩) + +@[simp] +lemma mk_real.apply (φ : {φ : K →+* ℂ // complex_embedding.is_real φ}) (x : K) : + complex.abs (φ x) = mk_real K φ x := apply φ x + +@[simp] +lemma mk_complex.apply (φ : {φ : K →+* ℂ // ¬ complex_embedding.is_real φ}) (x : K) : + complex.abs (φ x) = mk_complex K φ x := apply φ x + +variable [number_field K] + +lemma mk_complex.filter (w : { w : infinite_place K // w.is_complex }) : + finset.univ.filter (λ φ, mk_complex K φ = w) = + { ⟨w.1.embedding, is_complex_iff.1 w.2⟩, + ⟨complex_embedding.conjugate w.1.embedding, + complex_embedding.is_real_conjugate_iff.not.2 (is_complex_iff.1 w.2)⟩ } := +begin + ext φ, + simp_rw [finset.mem_filter, subtype.val_eq_coe, finset.mem_insert, finset.mem_singleton, + @subtype.ext_iff_val (infinite_place K), @subtype.ext_iff_val (K →+* ℂ), @eq_comm _ φ.val, + ← mk_eq_iff, mk_embedding, @eq_comm _ _ w.val], + simpa only [finset.mem_univ, true_and], +end + +lemma mk_complex.filter_card (w : { w : infinite_place K // w.is_complex }) : + (finset.univ.filter (λ φ, mk_complex K φ = w)).card = 2 := +begin + rw mk_complex.filter, + exact finset.card_doubleton + (subtype.mk_eq_mk.not.2 $ ne_comm.1 $ + complex_embedding.is_real_iff.not.1 $ is_complex_iff.1 w.2), +end + +noncomputable instance number_field.infinite_place.fintype : fintype (infinite_place K) := +set.fintype_range _ + +/-- The infinite part of the product formula : for `x ∈ K`, we have `Π_w ‖x‖_w = |norm(x)|` where +`‖·‖_w` is the normalized absolute value for `w`. -/ +lemma prod_eq_abs_norm (x : K) : + finset.univ.prod (λ w : infinite_place K, ite (w.is_real) (w x) ((w x) ^ 2)) = + abs (algebra.norm ℚ x) := +begin + convert (congr_arg complex.abs (@algebra.norm_eq_prod_embeddings ℚ _ _ _ _ ℂ _ _ _ _ _ x)).symm, + { rw [map_prod, ← equiv.prod_comp' ring_hom.equiv_rat_alg_hom (λ f, complex.abs (f x)) + (λ φ, complex.abs (φ x)) (λ _, by simpa only [ring_hom.equiv_rat_alg_hom_apply])], + dsimp only, + conv { to_rhs, congr, skip, funext, + rw ( by simp only [if_t_t] : complex.abs (f x) = + ite (complex_embedding.is_real f) (complex.abs (f x)) (complex.abs (f x))) }, + rw [finset.prod_ite, finset.prod_ite], + refine congr (congr_arg has_mul.mul _) _, + { rw [← finset.prod_subtype_eq_prod_filter, ← finset.prod_subtype_eq_prod_filter], + convert (equiv.prod_comp' (mk_real K) (λ φ, complex.abs (φ x)) (λ w, w x) _).symm, + any_goals { ext, simp only [finset.mem_subtype, finset.mem_univ], }, + exact λ φ, mk_real.apply K φ x, }, + { rw [finset.filter_congr (λ (w : infinite_place K) _, @not_is_real_iff_is_complex K _ w), + ← finset.prod_subtype_eq_prod_filter, ← finset.prod_subtype_eq_prod_filter], + convert finset.prod_fiberwise finset.univ (λ φ, mk_complex K φ) (λ φ, complex.abs (φ x)), + any_goals + { ext, simp only [finset.mem_subtype, finset.mem_univ, not_is_real_iff_is_complex], }, + { ext w, + rw [@finset.prod_congr _ _ _ _ _ (λ φ, w x) _ (eq.refl _) + (λ φ hφ, (mk_complex.apply K φ x).trans + (congr_fun (congr_arg coe_fn (finset.mem_filter.1 hφ).2) x)), finset.prod_const, + mk_complex.filter_card K w], + refl, }}}, + { rw [eq_rat_cast, ← complex.abs_of_real, complex.of_real_rat_cast], }, +end + +open fintype + +lemma card_real_embeddings : + card {φ : K →+* ℂ // complex_embedding.is_real φ} = card {w : infinite_place K // is_real w} := +by convert (fintype.of_equiv_card (mk_real K)).symm + +lemma card_complex_embeddings : + card {φ : K →+* ℂ // ¬ complex_embedding.is_real φ} = + 2 * card {w : infinite_place K // is_complex w} := +begin + rw [fintype.card, fintype.card, mul_comm, ← algebra.id.smul_eq_mul, ← finset.sum_const], + conv { to_rhs, congr, skip, funext, rw ← mk_complex.filter_card K x }, + simp_rw finset.card_eq_sum_ones, + exact (finset.sum_fiberwise finset.univ (λ φ, mk_complex K φ) (λ φ, 1)).symm +end + end number_field.infinite_place end infinite_place From c4bcf3ac336d52d09dfdfb295ecf234a56069edc Mon Sep 17 00:00:00 2001 From: leanprover-community-bot Date: Sat, 4 Feb 2023 03:54:23 +0000 Subject: [PATCH 0426/1291] chore(scripts): update nolints.txt (#18385) I am happy to remove some nolints for you! --- scripts/nolints.txt | 5 +---- scripts/style-exceptions.txt | 1 - 2 files changed, 1 insertion(+), 5 deletions(-) diff --git a/scripts/nolints.txt b/scripts/nolints.txt index a289f7daf787f..a2a306eb8dfd1 100644 --- a/scripts/nolints.txt +++ b/scripts/nolints.txt @@ -870,7 +870,4 @@ apply_nolint tactic.tidy.run_tactics doc_blame -- tactic/transfer.lean apply_nolint tactic.transfer doc_blame apply_nolint transfer.analyse_decls doc_blame -apply_nolint transfer.compute_transfer doc_blame - --- tactic/wlog.lean -apply_nolint tactic.wlog doc_blame \ No newline at end of file +apply_nolint transfer.compute_transfer doc_blame \ No newline at end of file diff --git a/scripts/style-exceptions.txt b/scripts/style-exceptions.txt index cb11f51c2c2a7..75be83d955367 100644 --- a/scripts/style-exceptions.txt +++ b/scripts/style-exceptions.txt @@ -73,4 +73,3 @@ src/tactic/tidy.lean : line 10 : ERR_MOD : Module docstring missing, or too late src/tactic/transfer.lean : line 6 : ERR_MOD : Module docstring missing, or too late src/tactic/transform_decl.lean : line 8 : ERR_MOD : Module docstring missing, or too late src/tactic/trunc_cases.lean : line 9 : ERR_MOD : Module docstring missing, or too late -src/tactic/wlog.lean : line 10 : ERR_MOD : Module docstring missing, or too late From 4c19a16e4b705bf135cf9a80ac18fcc99c438514 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Sat, 4 Feb 2023 09:12:11 +0000 Subject: [PATCH 0427/1291] chore(*): add mathlib4 synchronization comments (#18381) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `combinatorics.set_family.compression.uv` * `combinatorics.young.young_diagram` * `data.finmap` * `data.finsupp.multiset` * `data.finsupp.to_dfinsupp` * `data.fp.basic` * `data.set_like.fintype` * `dynamics.fixed_points.topology` * `linear_algebra.general_linear_group` * `order.countable_dense_linear_order` * `order.hom.lattice` * `order.succ_pred.linear_locally_finite` * `topology.omega_complete_partial_order` * `topology.uniform_space.cauchy` * `topology.uniform_space.separation` --- src/combinatorics/set_family/compression/uv.lean | 3 +++ src/combinatorics/young/young_diagram.lean | 3 +++ src/data/finmap.lean | 3 +++ src/data/finsupp/multiset.lean | 3 +++ src/data/finsupp/to_dfinsupp.lean | 3 +++ src/data/fp/basic.lean | 3 +++ src/data/set_like/fintype.lean | 3 +++ src/dynamics/fixed_points/topology.lean | 3 +++ src/linear_algebra/general_linear_group.lean | 3 +++ src/order/countable_dense_linear_order.lean | 3 +++ src/order/hom/lattice.lean | 3 +++ src/order/succ_pred/linear_locally_finite.lean | 3 +++ src/topology/omega_complete_partial_order.lean | 3 +++ src/topology/uniform_space/cauchy.lean | 3 +++ src/topology/uniform_space/separation.lean | 3 +++ 15 files changed, 45 insertions(+) diff --git a/src/combinatorics/set_family/compression/uv.lean b/src/combinatorics/set_family/compression/uv.lean index cdd5f6220c2e1..cec8c3800f4bb 100644 --- a/src/combinatorics/set_family/compression/uv.lean +++ b/src/combinatorics/set_family/compression/uv.lean @@ -8,6 +8,9 @@ import data.finset.card /-! # UV-compressions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines UV-compression. It is an operation on a set family that reduces its shadow. UV-compressing `a : α` along `u v : α` means replacing `a` by `(a ⊔ u) \ v` if `a` and `u` are diff --git a/src/combinatorics/young/young_diagram.lean b/src/combinatorics/young/young_diagram.lean index ed1174f80822d..b24cc54c3edab 100644 --- a/src/combinatorics/young/young_diagram.lean +++ b/src/combinatorics/young/young_diagram.lean @@ -9,6 +9,9 @@ import data.finset.preimage /-! # Young diagrams +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A Young diagram is a finite set of up-left justified boxes: ```text diff --git a/src/data/finmap.lean b/src/data/finmap.lean index 804f9bbb5ce41..2fb552f8cfdf3 100644 --- a/src/data/finmap.lean +++ b/src/data/finmap.lean @@ -8,6 +8,9 @@ import data.finset.basic import data.part /-! # Finite maps over `multiset` + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ universes u v w diff --git a/src/data/finsupp/multiset.lean b/src/data/finsupp/multiset.lean index 924aaa4ca80eb..f9d687b13f87f 100644 --- a/src/data/finsupp/multiset.lean +++ b/src/data/finsupp/multiset.lean @@ -9,6 +9,9 @@ import data.finsupp.order /-! # Equivalence between `multiset` and `ℕ`-valued finitely supported functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This defines `finsupp.to_multiset` the equivalence between `α →₀ ℕ` and `multiset α`, along with `multiset.to_finsupp` the reverse equivalence and `finsupp.order_iso_multiset` the equivalence promoted to an order isomorphism. diff --git a/src/data/finsupp/to_dfinsupp.lean b/src/data/finsupp/to_dfinsupp.lean index e4b018ab42d15..e855ec369bbb3 100644 --- a/src/data/finsupp/to_dfinsupp.lean +++ b/src/data/finsupp/to_dfinsupp.lean @@ -10,6 +10,9 @@ import data.finsupp.basic /-! # Conversion between `finsupp` and homogenous `dfinsupp` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This module provides conversions between `finsupp` and `dfinsupp`. It is in its own file since neither `finsupp` or `dfinsupp` depend on each other. diff --git a/src/data/fp/basic.lean b/src/data/fp/basic.lean index 4bbed9f2e17f8..3936873bc42a4 100644 --- a/src/data/fp/basic.lean +++ b/src/data/fp/basic.lean @@ -7,6 +7,9 @@ import data.semiquot import data.rat.floor /-! # Implementation of floating-point numbers (experimental). + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ def int.shift2 (a b : ℕ) : ℤ → ℕ × ℕ diff --git a/src/data/set_like/fintype.lean b/src/data/set_like/fintype.lean index 56d9a17466a3f..d1bb48f11a783 100644 --- a/src/data/set_like/fintype.lean +++ b/src/data/set_like/fintype.lean @@ -8,6 +8,9 @@ import data.fintype.powerset /-! # Set-like fintype +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains a fintype instance for set-like objects such as subgroups. If `set_like A B` and `fintype B` then `fintype A`. -/ diff --git a/src/dynamics/fixed_points/topology.lean b/src/dynamics/fixed_points/topology.lean index 88ce01ec5c485..ace366252c332 100644 --- a/src/dynamics/fixed_points/topology.lean +++ b/src/dynamics/fixed_points/topology.lean @@ -9,6 +9,9 @@ import topology.separation /-! # Topological properties of fixed points +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Currently this file contains two lemmas: - `is_fixed_pt_of_tendsto_iterate`: if `f^n(x) → y` and `f` is continuous at `y`, then `f y = y`; diff --git a/src/linear_algebra/general_linear_group.lean b/src/linear_algebra/general_linear_group.lean index d392c759b6dd0..f4629219b1ec4 100644 --- a/src/linear_algebra/general_linear_group.lean +++ b/src/linear_algebra/general_linear_group.lean @@ -8,6 +8,9 @@ import algebra.module.equiv /-! # The general linear group of linear maps +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The general linear group is defined to be the group of invertible linear maps from `M` to itself. See also `matrix.general_linear_group` diff --git a/src/order/countable_dense_linear_order.lean b/src/order/countable_dense_linear_order.lean index 36ec280490d45..f5deec2b668fb 100644 --- a/src/order/countable_dense_linear_order.lean +++ b/src/order/countable_dense_linear_order.lean @@ -9,6 +9,9 @@ import data.finset.lattice /-! # The back and forth method and countable dense linear orders +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Results Suppose `α β` are linear orders, with `α` countable and `β` dense, nontrivial. Then there is an diff --git a/src/order/hom/lattice.lean b/src/order/hom/lattice.lean index 033a64c543d15..7bda628e94690 100644 --- a/src/order/hom/lattice.lean +++ b/src/order/hom/lattice.lean @@ -10,6 +10,9 @@ import order.symm_diff /-! # Lattice homomorphisms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines (bounded) lattice homomorphisms. We use the `fun_like` design, so each type of morphisms has a companion typeclass which is meant to diff --git a/src/order/succ_pred/linear_locally_finite.lean b/src/order/succ_pred/linear_locally_finite.lean index 39db9c91883a8..2f0c15a1e5ba9 100644 --- a/src/order/succ_pred/linear_locally_finite.lean +++ b/src/order/succ_pred/linear_locally_finite.lean @@ -13,6 +13,9 @@ import logic.encodable.basic /-! # Linear locally finite orders +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We prove that a `linear_order` which is a `locally_finite_order` also verifies * `succ_order` * `pred_order` diff --git a/src/topology/omega_complete_partial_order.lean b/src/topology/omega_complete_partial_order.lean index b85acf2d4ae6d..785f8bcaf346b 100644 --- a/src/topology/omega_complete_partial_order.lean +++ b/src/topology/omega_complete_partial_order.lean @@ -9,6 +9,9 @@ import order.omega_complete_partial_order /-! # Scott Topological Spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A type of topological spaces whose notion of continuity is equivalent to continuity in ωCPOs. diff --git a/src/topology/uniform_space/cauchy.lean b/src/topology/uniform_space/cauchy.lean index 5425a039488cb..b96f8ef30b9b6 100644 --- a/src/topology/uniform_space/cauchy.lean +++ b/src/topology/uniform_space/cauchy.lean @@ -7,6 +7,9 @@ import topology.bases import topology.uniform_space.basic /-! # Theory of Cauchy filters in uniform spaces. Complete uniform spaces. Totally bounded subsets. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ universes u v diff --git a/src/topology/uniform_space/separation.lean b/src/topology/uniform_space/separation.lean index 5f9cdf39ce5d2..54c7cfe6ded99 100644 --- a/src/topology/uniform_space/separation.lean +++ b/src/topology/uniform_space/separation.lean @@ -11,6 +11,9 @@ import topology.separation /-! # Hausdorff properties of uniform spaces. Separation quotient. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file studies uniform spaces whose underlying topological spaces are separated (also known as Hausdorff or T₂). This turns out to be equivalent to asking that the intersection of all entourages From cc8e88c7c8c7bc80f91f84d11adb584bf9bd658f Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Sun, 5 Feb 2023 20:22:06 +0000 Subject: [PATCH 0428/1291] docs(algebra/module/linear_map): Fix a copy/paste error (#18389) --- src/algebra/module/linear_map.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/algebra/module/linear_map.lean b/src/algebra/module/linear_map.lean index ea7dfe2aa9095..17767837fbf4f 100644 --- a/src/algebra/module/linear_map.lean +++ b/src/algebra/module/linear_map.lean @@ -708,7 +708,7 @@ include σ₁₃ ext $ λ _, g.map_neg _ omit σ₁₃ -/-- The negation of a linear map is linear. -/ +/-- The subtraction of two linear maps is linear. -/ instance : has_sub (M →ₛₗ[σ₁₂] N₂) := ⟨λ f g, { to_fun := f - g, map_add' := λ x y, by simp only [pi.sub_apply, map_add, add_sub_add_comm], From 0a0ec35061ed9960bf0e7ffb0335f44447b58977 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Mon, 6 Feb 2023 06:28:11 +0000 Subject: [PATCH 0429/1291] chore(*): add mathlib4 synchronization comments (#18387) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `combinatorics.colex` * `combinatorics.set_family.harris_kleitman` * `data.finset.pimage` * `data.fintype.perm` * `data.holor` * `dynamics.minimal` * `order.filter.partial` * `order.partition.equipartition` * `order.partition.finpartition` * `topology.extend_from` * `topology.order.basic` * `topology.paracompact` * `topology.shrinking_lemma` * `topology.uniform_space.absolute_value` * `topology.uniform_space.complete_separated` * `topology.uniform_space.pi` * `topology.uniform_space.uniform_convergence` * `topology.uniform_space.uniform_embedding` --- src/combinatorics/colex.lean | 3 +++ src/combinatorics/set_family/harris_kleitman.lean | 3 +++ src/data/finset/pimage.lean | 3 +++ src/data/fintype/perm.lean | 3 +++ src/data/holor.lean | 3 +++ src/dynamics/minimal.lean | 3 +++ src/order/filter/partial.lean | 3 +++ src/order/partition/equipartition.lean | 3 +++ src/order/partition/finpartition.lean | 3 +++ src/topology/extend_from.lean | 3 +++ src/topology/order/basic.lean | 3 +++ src/topology/paracompact.lean | 3 +++ src/topology/shrinking_lemma.lean | 3 +++ src/topology/uniform_space/absolute_value.lean | 3 +++ src/topology/uniform_space/complete_separated.lean | 3 +++ src/topology/uniform_space/pi.lean | 3 +++ src/topology/uniform_space/uniform_convergence.lean | 3 +++ src/topology/uniform_space/uniform_embedding.lean | 3 +++ 18 files changed, 54 insertions(+) diff --git a/src/combinatorics/colex.lean b/src/combinatorics/colex.lean index 0441534ef83ca..cc90e329aa314 100644 --- a/src/combinatorics/colex.lean +++ b/src/combinatorics/colex.lean @@ -9,6 +9,9 @@ import algebra.geom_sum /-! # Colex +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define the colex ordering for finite sets, and give a couple of important lemmas and properties relating to it. diff --git a/src/combinatorics/set_family/harris_kleitman.lean b/src/combinatorics/set_family/harris_kleitman.lean index c4654dd7dc7a0..236ffca1e72bf 100644 --- a/src/combinatorics/set_family/harris_kleitman.lean +++ b/src/combinatorics/set_family/harris_kleitman.lean @@ -10,6 +10,9 @@ import data.fintype.big_operators /-! # Harris-Kleitman inequality +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves the Harris-Kleitman inequality. This relates `𝒜.card * ℬ.card` and `2 ^ card α * (𝒜 ∩ ℬ).card` where `𝒜` and `ℬ` are upward- or downcard-closed finite families of finsets. This can be interpreted as saying that any two lower sets (resp. any two upper sets) diff --git a/src/data/finset/pimage.lean b/src/data/finset/pimage.lean index 7b2a3f6aec4bb..9ed110e3d7479 100644 --- a/src/data/finset/pimage.lean +++ b/src/data/finset/pimage.lean @@ -9,6 +9,9 @@ import data.pfun /-! # Image of a `finset α` under a partially defined function +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define `part.to_finset` and `finset.pimage`. We also prove some trivial lemmas about these definitions. diff --git a/src/data/fintype/perm.lean b/src/data/fintype/perm.lean index a535e13efeee7..c3609830054bd 100644 --- a/src/data/fintype/perm.lean +++ b/src/data/fintype/perm.lean @@ -9,6 +9,9 @@ import group_theory.perm.basic /-! # fintype instances for `equiv` and `perm` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Main declarations: * `perms_of_finset s`: The finset of permutations of the finset `s`. diff --git a/src/data/holor.lean b/src/data/holor.lean index 86e226a66904f..18a484aba6560 100644 --- a/src/data/holor.lean +++ b/src/data/holor.lean @@ -9,6 +9,9 @@ import algebra.big_operators.basic /-! # Basic properties of holors +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Holors are indexed collections of tensor coefficients. Confusingly, they are often called tensors in physics and in the neural network community. diff --git a/src/dynamics/minimal.lean b/src/dynamics/minimal.lean index 73c2624ad51b1..0f2dbe871cb8d 100644 --- a/src/dynamics/minimal.lean +++ b/src/dynamics/minimal.lean @@ -9,6 +9,9 @@ import topology.algebra.const_mul_action /-! # Minimal action of a group +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define an action of a monoid `M` on a topological space `α` to be *minimal* if the `M`-orbit of every point `x : α` is dense. We also provide an additive version of this definition and prove some basic facts about minimal actions. diff --git a/src/order/filter/partial.lean b/src/order/filter/partial.lean index fd96d8812e966..2abe4b83c2d33 100644 --- a/src/order/filter/partial.lean +++ b/src/order/filter/partial.lean @@ -9,6 +9,9 @@ import data.pfun /-! # `tendsto` for relations and partial functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file generalizes `filter` definitions from functions to partial functions and relations. ## Considering functions and partial functions as relations diff --git a/src/order/partition/equipartition.lean b/src/order/partition/equipartition.lean index 3a3d8b1f6eddb..1a733509f0335 100644 --- a/src/order/partition/equipartition.lean +++ b/src/order/partition/equipartition.lean @@ -9,6 +9,9 @@ import order.partition.finpartition /-! # Finite equipartitions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines finite equipartitions, the partitions whose parts all are the same size up to a difference of `1`. diff --git a/src/order/partition/finpartition.lean b/src/order/partition/finpartition.lean index c20d35dafaaaf..77687a71a5f46 100644 --- a/src/order/partition/finpartition.lean +++ b/src/order/partition/finpartition.lean @@ -10,6 +10,9 @@ import order.sup_indep /-! # Finite partitions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we define finite partitions. A finpartition of `a : α` is a finite set of pairwise disjoint parts `parts : finset α` which does not contain `⊥` and whose supremum is `a`. diff --git a/src/topology/extend_from.lean b/src/topology/extend_from.lean index 23838fd18eb1f..5fa21fb42a4d0 100644 --- a/src/topology/extend_from.lean +++ b/src/topology/extend_from.lean @@ -8,6 +8,9 @@ import topology.separation /-! # Extending a function from a subset +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The main definition of this file is `extend_from A f` where `f : X → Y` and `A : set X`. This defines a new function `g : X → Y` which maps any `x₀ : X` to the limit of `f` as `x` tends to `x₀`, if such a limit exists. diff --git a/src/topology/order/basic.lean b/src/topology/order/basic.lean index d7a7813ed49f0..51e9d3489e4f6 100644 --- a/src/topology/order/basic.lean +++ b/src/topology/order/basic.lean @@ -12,6 +12,9 @@ import topology.algebra.order.left_right /-! # Theory of topology on ordered spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions The order topology on an ordered space is the topology generated by all open intervals (or diff --git a/src/topology/paracompact.lean b/src/topology/paracompact.lean index 505621d65b45e..672ffe7f71075 100644 --- a/src/topology/paracompact.lean +++ b/src/topology/paracompact.lean @@ -10,6 +10,9 @@ import data.option.basic /-! # Paracompact topological spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A topological space `X` is said to be paracompact if every open covering of `X` admits a locally finite refinement. diff --git a/src/topology/shrinking_lemma.lean b/src/topology/shrinking_lemma.lean index 94ca1524a6844..0212e4856a501 100644 --- a/src/topology/shrinking_lemma.lean +++ b/src/topology/shrinking_lemma.lean @@ -8,6 +8,9 @@ import topology.separation /-! # The shrinking lemma +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove a few versions of the shrinking lemma. The lemma says that in a normal topological space a point finite open covering can be “shrunk”: for a point finite open covering `u : ι → set X` there exists a refinement `v : ι → set X` such that `closure (v i) ⊆ u i`. diff --git a/src/topology/uniform_space/absolute_value.lean b/src/topology/uniform_space/absolute_value.lean index 42e7a80431858..5796516c80e8e 100644 --- a/src/topology/uniform_space/absolute_value.lean +++ b/src/topology/uniform_space/absolute_value.lean @@ -9,6 +9,9 @@ import topology.uniform_space.basic /-! # Uniform structure induced by an absolute value +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We build a uniform space structure on a commutative ring `R` equipped with an absolute value into a linear ordered field `𝕜`. Of course in the case `R` is `ℚ`, `ℝ` or `ℂ` and `𝕜 = ℝ`, we get the same thing as the metric space construction, and the general construction diff --git a/src/topology/uniform_space/complete_separated.lean b/src/topology/uniform_space/complete_separated.lean index da9939d6a9848..8e5295cc88668 100644 --- a/src/topology/uniform_space/complete_separated.lean +++ b/src/topology/uniform_space/complete_separated.lean @@ -10,6 +10,9 @@ import topology.dense_embedding /-! # Theory of complete separated uniform spaces. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file is for elementary lemmas that depend on both Cauchy filters and separation. -/ diff --git a/src/topology/uniform_space/pi.lean b/src/topology/uniform_space/pi.lean index 12ca6fde6a6a9..7f7880e29dad1 100644 --- a/src/topology/uniform_space/pi.lean +++ b/src/topology/uniform_space/pi.lean @@ -8,6 +8,9 @@ import topology.uniform_space.separation /-! # Indexed product of uniform spaces + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ noncomputable theory diff --git a/src/topology/uniform_space/uniform_convergence.lean b/src/topology/uniform_space/uniform_convergence.lean index a1e146c7de013..04ea04de1256c 100644 --- a/src/topology/uniform_space/uniform_convergence.lean +++ b/src/topology/uniform_space/uniform_convergence.lean @@ -10,6 +10,9 @@ import topology.uniform_space.cauchy /-! # Uniform convergence +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A sequence of functions `Fₙ` (with values in a metric space) converges uniformly on a set `s` to a function `f` if, for all `ε > 0`, for all large enough `n`, one has for all `y ∈ s` the inequality `dist (f y, Fₙ y) < ε`. Under uniform convergence, many properties of the `Fₙ` pass to the limit, diff --git a/src/topology/uniform_space/uniform_embedding.lean b/src/topology/uniform_space/uniform_embedding.lean index 261ab5583f2f6..5f3e92c429039 100644 --- a/src/topology/uniform_space/uniform_embedding.lean +++ b/src/topology/uniform_space/uniform_embedding.lean @@ -10,6 +10,9 @@ import topology.dense_embedding /-! # Uniform embeddings of uniform spaces. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Extension of uniform continuous functions. -/ From bb37dbda903641effc74366a2774cefdf2c6734d Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 6 Feb 2023 23:03:15 +0000 Subject: [PATCH 0430/1291] feat(algebra/big_operators/order): prod_pos lemma for `ennreal` (#18391) This adds: * `canonically_ordered_comm_semiring.list_prod_pos` * `canonically_ordered_comm_semiring.multiset_prod_pos` * `canonically_ordered_comm_semiring.prod_pos` which extend the existing `canonically_ordered_comm_semiring.mul_pos`. Primarily, these are intended for use on `enat` and `ennreal`, which don't satisfy the typeclasses required by `list.prod_pos` and `finset.prod_pos`. At any rate, those statements are weaker. Forward port in https://github.com/leanprover-community/mathlib4/pull/2120 --- src/algebra/big_operators/order.lean | 15 +++++++++++++++ src/data/list/big_operators/basic.lean | 10 ++++++++++ 2 files changed, 25 insertions(+) diff --git a/src/algebra/big_operators/order.lean b/src/algebra/big_operators/order.lean index 413f7f89035dc..ac17f0708b860 100644 --- a/src/algebra/big_operators/order.lean +++ b/src/algebra/big_operators/order.lean @@ -536,6 +536,21 @@ section canonically_ordered_comm_semiring variables [canonically_ordered_comm_semiring R] {f g h : ι → R} {s : finset ι} {i : ι} +@[simp] +lemma _root_.canonically_ordered_comm_semiring.multiset_prod_pos [nontrivial R] {m : multiset R} : + 0 < m.prod ↔ (∀ x ∈ m, (0 : R) < x) := +begin + induction m using quotient.induction_on, + rw [multiset.quot_mk_to_coe, multiset.coe_prod], + exact canonically_ordered_comm_semiring.list_prod_pos, +end + +/-- Note that the name is to match `canonically_ordered_comm_semiring.mul_pos`. -/ +@[simp] +lemma _root_.canonically_ordered_comm_semiring.prod_pos [nontrivial R] : + 0 < ∏ i in s, f i ↔ (∀ i ∈ s, (0 : R) < f i) := +canonically_ordered_comm_semiring.multiset_prod_pos.trans $ by simp + lemma prod_le_prod' (h : ∀ i ∈ s, f i ≤ g i) : ∏ i in s, f i ≤ ∏ i in s, g i := begin diff --git a/src/data/list/big_operators/basic.lean b/src/data/list/big_operators/basic.lean index b49bb59ddf5fa..c93af87dad977 100644 --- a/src/data/list/big_operators/basic.lean +++ b/src/data/list/big_operators/basic.lean @@ -486,6 +486,16 @@ begin exact mul_pos (h _ $ mem_cons_self _ _) (ih $ λ a ha, h a $ mem_cons_of_mem _ ha) } end +/-- A variant of `list.prod_pos` for `canonically_ordered_comm_semiring`. -/ +@[simp] +lemma _root_.canonically_ordered_comm_semiring.list_prod_pos + {α : Type*} [canonically_ordered_comm_semiring α] [nontrivial α] : + Π {l : list α}, 0 < l.prod ↔ (∀ x ∈ l, (0 : α) < x) +| [] := ⟨λ h x hx, hx.elim, λ _, zero_lt_one⟩ +| (x :: xs) := by simp_rw [prod_cons, mem_cons_iff, forall_eq_or_imp, + canonically_ordered_comm_semiring.mul_pos, + _root_.canonically_ordered_comm_semiring.list_prod_pos] + /-! Several lemmas about sum/head/tail for `list ℕ`. These are hard to generalize well, as they rely on the fact that `default ℕ = 0`. From 2d6f88c296da8df484d7f5b9ee1d10910ab473a2 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 7 Feb 2023 01:58:42 +0000 Subject: [PATCH 0431/1291] chore: Add toml file for archive and counterexamples (#18388) The end goal here is to surface these in doc-gen. For doc-gen to behave these need to be mutually importable. A step is added to the "test" run that verifies this by trying to import everything at the same time. For this new test to pass, some files need namespaces wrapped around their content. Probably we should be more principled about this namespace wrapping, but that can be left as future work, as all we really care about to start integrating with doc-gen is avoiding clashes. --- .github/workflows/bors.yml | 15 ++++++++++++++- .github/workflows/build.yml | 15 ++++++++++++++- .github/workflows/build.yml.in | 15 ++++++++++++++- .github/workflows/build_fork.yml | 15 ++++++++++++++- .gitignore | 2 +- archive/100-theorems-list/82_cubing_a_cube.lean | 5 ++++- archive/imo/imo1960_q1.lean | 6 ++++++ archive/imo/imo1962_q1.lean | 6 ++++++ archive/imo/imo1964_q1.lean | 6 ++++++ archive/imo/imo1969_q1.lean | 6 ++++++ archive/imo/imo1981_q3.lean | 6 ++++-- archive/leanpkg.toml | 7 +++++++ counterexamples/homogeneous_prime_not_prime.lean | 2 +- counterexamples/leanpkg.toml | 7 +++++++ .../linear_order_with_pos_mul_pos_eq_zero.lean | 2 +- counterexamples/map_floor.lean | 4 ++-- 16 files changed, 107 insertions(+), 12 deletions(-) create mode 100644 archive/leanpkg.toml create mode 100644 counterexamples/leanpkg.toml diff --git a/.github/workflows/bors.yml b/.github/workflows/bors.yml index e43db8caf1548..e808ac75f1a2e 100644 --- a/.github/workflows/bors.yml +++ b/.github/workflows/bors.yml @@ -192,13 +192,26 @@ jobs: - name: install Python dependencies if: ${{ ! 1 }} - run: python3 -m pip install --upgrade pip pyyaml requests + run: python3 -m pip install --upgrade pip pyyaml requests mathlibtools - name: tests run: | set -o pipefail lean --json -T100000 --make docs archive roadmap test counterexamples | python3 scripts/detect_errors.py + - name: check archive and counterexample directories have unique identifiers + run: | + pip install mathlibtools + (cd archive && python -m mathlibtools.leanproject mk-all && mv all.lean archive_all.lean) + (cd counterexamples && python -m mathlibtools.leanproject mk-all && mv all.lean counterexamples_all.lean) + python -m mathlibtools.leanproject mk-all + echo "import all" >> all_all.lean + echo "import counterexamples_all" >> all_all.lean + echo "import archive_all" >> all_all.lean + echo "path ./archive" >> leanpkg.path + echo "path ./counterexamples" >> leanpkg.path + lean --json -T100000 --make all_all.lean | python3 scripts/detect_errors.py + - uses: actions/setup-java@v2 if: ${{ ! 1 }} with: diff --git a/.github/workflows/build.yml b/.github/workflows/build.yml index 8296a49e4e11b..0656b61d919ae 100644 --- a/.github/workflows/build.yml +++ b/.github/workflows/build.yml @@ -200,13 +200,26 @@ jobs: - name: install Python dependencies if: ${{ ! 1 }} - run: python3 -m pip install --upgrade pip pyyaml requests + run: python3 -m pip install --upgrade pip pyyaml requests mathlibtools - name: tests run: | set -o pipefail lean --json -T100000 --make docs archive roadmap test counterexamples | python3 scripts/detect_errors.py + - name: check archive and counterexample directories have unique identifiers + run: | + pip install mathlibtools + (cd archive && python -m mathlibtools.leanproject mk-all && mv all.lean archive_all.lean) + (cd counterexamples && python -m mathlibtools.leanproject mk-all && mv all.lean counterexamples_all.lean) + python -m mathlibtools.leanproject mk-all + echo "import all" >> all_all.lean + echo "import counterexamples_all" >> all_all.lean + echo "import archive_all" >> all_all.lean + echo "path ./archive" >> leanpkg.path + echo "path ./counterexamples" >> leanpkg.path + lean --json -T100000 --make all_all.lean | python3 scripts/detect_errors.py + - uses: actions/setup-java@v2 if: ${{ ! 1 }} with: diff --git a/.github/workflows/build.yml.in b/.github/workflows/build.yml.in index ffed44267c589..f17250cc6b119 100644 --- a/.github/workflows/build.yml.in +++ b/.github/workflows/build.yml.in @@ -178,13 +178,26 @@ jobs: - name: install Python dependencies if: ${{ ! IS_SELF_HOSTED }} - run: python3 -m pip install --upgrade pip pyyaml requests + run: python3 -m pip install --upgrade pip pyyaml requests mathlibtools - name: tests run: | set -o pipefail lean --json -T100000 --make docs archive roadmap test counterexamples | python3 scripts/detect_errors.py + - name: check archive and counterexample directories have unique identifiers + run: | + pip install mathlibtools + (cd archive && python -m mathlibtools.leanproject mk-all && mv all.lean archive_all.lean) + (cd counterexamples && python -m mathlibtools.leanproject mk-all && mv all.lean counterexamples_all.lean) + python -m mathlibtools.leanproject mk-all + echo "import all" >> all_all.lean + echo "import counterexamples_all" >> all_all.lean + echo "import archive_all" >> all_all.lean + echo "path ./archive" >> leanpkg.path + echo "path ./counterexamples" >> leanpkg.path + lean --json -T100000 --make all_all.lean | python3 scripts/detect_errors.py + - uses: actions/setup-java@v2 if: ${{ ! IS_SELF_HOSTED }} with: diff --git a/.github/workflows/build_fork.yml b/.github/workflows/build_fork.yml index 3a093a8e654f4..8f12c3bc1f37d 100644 --- a/.github/workflows/build_fork.yml +++ b/.github/workflows/build_fork.yml @@ -198,13 +198,26 @@ jobs: - name: install Python dependencies if: ${{ ! 0 }} - run: python3 -m pip install --upgrade pip pyyaml requests + run: python3 -m pip install --upgrade pip pyyaml requests mathlibtools - name: tests run: | set -o pipefail lean --json -T100000 --make docs archive roadmap test counterexamples | python3 scripts/detect_errors.py + - name: check archive and counterexample directories have unique identifiers + run: | + pip install mathlibtools + (cd archive && python -m mathlibtools.leanproject mk-all && mv all.lean archive_all.lean) + (cd counterexamples && python -m mathlibtools.leanproject mk-all && mv all.lean counterexamples_all.lean) + python -m mathlibtools.leanproject mk-all + echo "import all" >> all_all.lean + echo "import counterexamples_all" >> all_all.lean + echo "import archive_all" >> all_all.lean + echo "path ./archive" >> leanpkg.path + echo "path ./counterexamples" >> leanpkg.path + lean --json -T100000 --make all_all.lean | python3 scripts/detect_errors.py + - uses: actions/setup-java@v2 if: ${{ ! 0 }} with: diff --git a/.gitignore b/.gitignore index ed583e63ad380..a7798b842bb80 100644 --- a/.gitignore +++ b/.gitignore @@ -1,6 +1,6 @@ *.olean /_target -/leanpkg.path +leanpkg.path _cache __pycache__ all.lean diff --git a/archive/100-theorems-list/82_cubing_a_cube.lean b/archive/100-theorems-list/82_cubing_a_cube.lean index 330a23188f703..ff662f64b69eb 100644 --- a/archive/100-theorems-list/82_cubing_a_cube.lean +++ b/archive/100-theorems-list/82_cubing_a_cube.lean @@ -16,11 +16,12 @@ We follow the proof described here: http://www.alaricstephen.com/main-featured/2017/9/28/cubing-a-cube-proof -/ - open real set function fin noncomputable theory +namespace «82» + variable {n : ℕ} /-- Given three intervals `I, J, K` such that `J ⊂ I`, @@ -517,3 +518,5 @@ begin exact @not_correct n s coe hfin.to_subtype h2.coe_sort ⟨hd.subtype _ _, (Union_subtype _ _).trans hU, hinj.injective, hn⟩ end + +end «82» diff --git a/archive/imo/imo1960_q1.lean b/archive/imo/imo1960_q1.lean index 7caef8294f29b..fa3ee7680092f 100644 --- a/archive/imo/imo1960_q1.lean +++ b/archive/imo/imo1960_q1.lean @@ -22,6 +22,8 @@ The strategy here is roughly brute force, checking the possible multiples of 11. open nat +namespace imo1960_q1 + def sum_of_squares (L : list ℕ) : ℕ := (L.map (λ x, x * x)).sum def problem_predicate (n : ℕ) : Prop := @@ -98,5 +100,9 @@ Now we just need to prove the equivalence, for the precise problem statement. lemma left_direction (n : ℕ) (spn : solution_predicate n) : problem_predicate n := by rcases spn with (rfl | rfl); norm_num [problem_predicate, sum_of_squares] +end imo1960_q1 + +open imo1960_q1 + theorem imo1960_q1 (n : ℕ) : problem_predicate n ↔ solution_predicate n := ⟨right_direction, left_direction n⟩ diff --git a/archive/imo/imo1962_q1.lean b/archive/imo/imo1962_q1.lean index 7bd2c6f015d7c..0afc5e6be6726 100644 --- a/archive/imo/imo1962_q1.lean +++ b/archive/imo/imo1962_q1.lean @@ -21,6 +21,8 @@ we define the problem as a predicate, and then prove a particular number is the of a set satisfying it. -/ +namespace imo1962_q1 + open nat def problem_predicate (n : ℕ) : Prop := @@ -154,5 +156,9 @@ begin exact h5.ge, }, }, end +end imo1962_q1 + +open imo1962_q1 + theorem imo1962_q1 : is_least {n | problem_predicate n} 153846 := ⟨satisfied_by_153846, no_smaller_solutions⟩ diff --git a/archive/imo/imo1964_q1.lean b/archive/imo/imo1964_q1.lean index ac90068a7afb6..fd987ba4f5f10 100644 --- a/archive/imo/imo1964_q1.lean +++ b/archive/imo/imo1964_q1.lean @@ -24,6 +24,8 @@ integers which are a multiple of 3. open nat +namespace imo1964_q1 + lemma two_pow_three_mul_mod_seven (m : ℕ) : 2 ^ (3 * m) ≡ 1 [MOD 7] := begin rw pow_mul, @@ -79,6 +81,10 @@ begin apply two_pow_three_mul_mod_seven } end +end imo1964_q1 + +open imo1964_q1 + theorem imo1964_q1b (n : ℕ) : ¬ (7 ∣ 2 ^ n + 1) := begin let t := n % 3, diff --git a/archive/imo/imo1969_q1.lean b/archive/imo/imo1969_q1.lean index a2ccb8cca33e6..3eee264d353ca 100644 --- a/archive/imo/imo1969_q1.lean +++ b/archive/imo/imo1969_q1.lean @@ -17,6 +17,8 @@ the number $z = n^4 + a$ is not prime for any natural number $n$. open int nat +namespace imo1969_q1 + /-- `good_nats` is the set of natural numbers satisfying the condition in the problem statement, namely the `a : ℕ` such that `n^4 + a` is not prime for any `n : ℕ`. -/ def good_nats : set ℕ := {a : ℕ | ∀ n : ℕ, ¬ nat.prime (n^4 + a)} @@ -67,6 +69,10 @@ in the `strict_mono` namespace. -/ lemma a_choice_strict_mono : strict_mono a_choice := ((strict_mono_id.const_add 2).nat_pow (dec_trivial : 0 < 4)).const_mul (dec_trivial : 0 < 4) +end imo1969_q1 + +open imo1969_q1 + /-- We conclude by using the fact that `a_choice` is an injective function from the natural numbers to the set `good_nats`. -/ theorem imo1969_q1 : set.infinite {a : ℕ | ∀ n : ℕ, ¬ nat.prime (n^4 + a)} := diff --git a/archive/imo/imo1981_q3.lean b/archive/imo/imo1981_q3.lean index ea34d76d03bf2..4b40612e5a873 100644 --- a/archive/imo/imo1981_q3.lean +++ b/archive/imo/imo1981_q3.lean @@ -24,7 +24,7 @@ We first generalize the problem to `{1, 2, ..., N}` and specialize to `N = 1981` -/ open int nat set -section +namespace imo1981_q3 variable (N : ℕ) -- N = 1981 @[mk_iff] structure problem_predicate (m n : ℤ) : Prop := @@ -189,7 +189,9 @@ theorem solution_greatest (H : problem_predicate N (fib K) (fib (K + 1))) : is_greatest (specified_set N) M := ⟨⟨fib K, fib (K+1), by simp [HM], H⟩, λ k h, solution_bound HK HM h⟩ -end +end imo1981_q3 + +open imo1981_q3 /- Now we just have to demonstrate that 987 and 1597 are in fact the largest Fibonacci diff --git a/archive/leanpkg.toml b/archive/leanpkg.toml new file mode 100644 index 0000000000000..a89c0e895b5f7 --- /dev/null +++ b/archive/leanpkg.toml @@ -0,0 +1,7 @@ +[package] +name = "mathlib-archive" +version = "0.1" +path = "." + +[dependencies] +mathlib = {path = ".."} diff --git a/counterexamples/homogeneous_prime_not_prime.lean b/counterexamples/homogeneous_prime_not_prime.lean index 8b9f1a9328952..8692774542313 100644 --- a/counterexamples/homogeneous_prime_not_prime.lean +++ b/counterexamples/homogeneous_prime_not_prime.lean @@ -74,7 +74,7 @@ lemma grading.mul_mem : ∀ ⦃i j : two⦄ {a b : (R × R)} (ha : a ∈ grading end -notation `R` := zmod 4 +local notation `R` := zmod 4 /-- `R² ≅ {(a, a) | a ∈ R} ⨁ {(0, b) | b ∈ R}` by `(x, y) ↦ (x, x) + (0, y - x)`. -/ def grading.decompose : (R × R) →+ direct_sum two (λ i, grading R i) := diff --git a/counterexamples/leanpkg.toml b/counterexamples/leanpkg.toml new file mode 100644 index 0000000000000..814d73c4e45c1 --- /dev/null +++ b/counterexamples/leanpkg.toml @@ -0,0 +1,7 @@ +[package] +name = "mathlib-counterexamples" +version = "0.1" +path = "." + +[dependencies] +mathlib = {path = ".."} diff --git a/counterexamples/linear_order_with_pos_mul_pos_eq_zero.lean b/counterexamples/linear_order_with_pos_mul_pos_eq_zero.lean index 28298a847d2cc..6e08c5f9f847f 100644 --- a/counterexamples/linear_order_with_pos_mul_pos_eq_zero.lean +++ b/counterexamples/linear_order_with_pos_mul_pos_eq_zero.lean @@ -32,7 +32,7 @@ instance inhabited : inhabited foo := ⟨zero⟩ instance : has_zero foo := ⟨zero⟩ instance : has_one foo := ⟨one⟩ -notation `ε` := eps +local notation `ε` := eps /-- The order on `foo` is the one induced by the natural order on the image of `aux1`. -/ def aux1 : foo → ℕ diff --git a/counterexamples/map_floor.lean b/counterexamples/map_floor.lean index 8c00769a84eaf..b1c262d6d780c 100644 --- a/counterexamples/map_floor.lean +++ b/counterexamples/map_floor.lean @@ -41,9 +41,9 @@ open_locale polynomial /-- The integers with infinitesimals adjoined. -/ @[derive [comm_ring, nontrivial, inhabited]] def int_with_epsilon := ℤ[X] -notation `ℤ[ε]` := int_with_epsilon +local notation `ℤ[ε]` := int_with_epsilon -notation `ε` := (X : ℤ[ε]) +local notation `ε` := (X : ℤ[ε]) namespace int_with_epsilon From 50832daea47b195a48b5b33b1c8b2162c48c3afc Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Tue, 7 Feb 2023 04:08:25 +0000 Subject: [PATCH 0432/1291] chore(*): add mathlib4 synchronization comments (#18393) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.char_zero.quotient` * `algebra.periodic` * `algebra.star.self_adjoint` * `combinatorics.additive.ruzsa_covering` * `combinatorics.set_family.kleitman` * `combinatorics.young.semistandard_tableau` * `data.dfinsupp.order` * `data.fin.tuple.bubble_sort_induction` * `data.multiset.fintype` * `data.nat.part_enat` * `data.pnat.factors` * `data.setoid.partition` * `data.sigma.interval` * `group_theory.double_coset` * `group_theory.perm.subgroup` * `measure_theory.measurable_space_def` * `order.heyting.hom` * `order.hom.complete_lattice` * `topology.G_delta` * `topology.algebra.order.archimedean` * `topology.algebra.order.compact` * `topology.algebra.order.extr_closure` * `topology.algebra.order.intermediate_value` * `topology.algebra.order.monotone_continuity` * `topology.algebra.order.monotone_convergence` * `topology.algebra.order.proj_Icc` * `topology.algebra.order.t5` * `topology.algebra.star` * `topology.continuous_function.basic` * `topology.instances.sign` * `topology.order.priestley` * `topology.perfect` --- src/algebra/char_zero/quotient.lean | 3 +++ src/algebra/periodic.lean | 3 +++ src/algebra/star/self_adjoint.lean | 3 +++ src/combinatorics/additive/ruzsa_covering.lean | 3 +++ src/combinatorics/set_family/kleitman.lean | 3 +++ src/combinatorics/young/semistandard_tableau.lean | 3 +++ src/data/dfinsupp/order.lean | 3 +++ src/data/fin/tuple/bubble_sort_induction.lean | 3 +++ src/data/multiset/fintype.lean | 3 +++ src/data/nat/part_enat.lean | 3 +++ src/data/pnat/factors.lean | 3 +++ src/data/setoid/partition.lean | 3 +++ src/data/sigma/interval.lean | 3 +++ src/group_theory/double_coset.lean | 3 +++ src/group_theory/perm/subgroup.lean | 3 +++ src/measure_theory/measurable_space_def.lean | 3 +++ src/order/heyting/hom.lean | 3 +++ src/order/hom/complete_lattice.lean | 3 +++ src/topology/G_delta.lean | 3 +++ src/topology/algebra/order/archimedean.lean | 3 +++ src/topology/algebra/order/compact.lean | 3 +++ src/topology/algebra/order/extr_closure.lean | 3 +++ src/topology/algebra/order/intermediate_value.lean | 3 +++ src/topology/algebra/order/monotone_continuity.lean | 3 +++ src/topology/algebra/order/monotone_convergence.lean | 3 +++ src/topology/algebra/order/proj_Icc.lean | 3 +++ src/topology/algebra/order/t5.lean | 3 +++ src/topology/algebra/star.lean | 3 +++ src/topology/continuous_function/basic.lean | 3 +++ src/topology/instances/sign.lean | 3 +++ src/topology/order/priestley.lean | 3 +++ src/topology/perfect.lean | 3 +++ 32 files changed, 96 insertions(+) diff --git a/src/algebra/char_zero/quotient.lean b/src/algebra/char_zero/quotient.lean index 0503face396f1..61ad45384b7d0 100644 --- a/src/algebra/char_zero/quotient.lean +++ b/src/algebra/char_zero/quotient.lean @@ -7,6 +7,9 @@ import group_theory.quotient_group /-! # Lemmas about quotients in characteristic zero + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ variables {R : Type*} [division_ring R] [char_zero R] {p : R} diff --git a/src/algebra/periodic.lean b/src/algebra/periodic.lean index c4a5570dd2f21..fba880b221155 100644 --- a/src/algebra/periodic.lean +++ b/src/algebra/periodic.lean @@ -15,6 +15,9 @@ import group_theory.submonoid.membership /-! # Periodicity +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define and then prove facts about periodic and antiperiodic functions. ## Main definitions diff --git a/src/algebra/star/self_adjoint.lean b/src/algebra/star/self_adjoint.lean index 81aa4b3834f10..4526e2bbf58cd 100644 --- a/src/algebra/star/self_adjoint.lean +++ b/src/algebra/star/self_adjoint.lean @@ -10,6 +10,9 @@ import group_theory.subgroup.basic /-! # Self-adjoint, skew-adjoint and normal elements of a star additive group +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines `self_adjoint R` (resp. `skew_adjoint R`), where `R` is a star additive group, as the additive subgroup containing the elements that satisfy `star x = x` (resp. `star x = -x`). This includes, for instance, (skew-)Hermitian operators on Hilbert spaces. diff --git a/src/combinatorics/additive/ruzsa_covering.lean b/src/combinatorics/additive/ruzsa_covering.lean index 56ffc3abe0e5f..a7f65da9f18fa 100644 --- a/src/combinatorics/additive/ruzsa_covering.lean +++ b/src/combinatorics/additive/ruzsa_covering.lean @@ -8,6 +8,9 @@ import data.finset.pointwise /-! # Ruzsa's covering lemma +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves the Ruzsa covering lemma. This says that, for `s`, `t` finsets, we can cover `s` with at most `(s + t).card / t.card` copies of `t - t`. diff --git a/src/combinatorics/set_family/kleitman.lean b/src/combinatorics/set_family/kleitman.lean index 3fe9a533bc0c8..4d1edd2943526 100644 --- a/src/combinatorics/set_family/kleitman.lean +++ b/src/combinatorics/set_family/kleitman.lean @@ -9,6 +9,9 @@ import combinatorics.set_family.intersecting /-! # Kleitman's bound on the size of intersecting families +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + An intersecting family on `n` elements has size at most `2ⁿ⁻¹`, so we could naïvely think that two intersecting families could cover all `2ⁿ` sets. But actually that's not case because for example none of them can contain the empty set. Intersecting families are in some sense correlated. diff --git a/src/combinatorics/young/semistandard_tableau.lean b/src/combinatorics/young/semistandard_tableau.lean index 993b41bef0539..4c1261a411e9d 100644 --- a/src/combinatorics/young/semistandard_tableau.lean +++ b/src/combinatorics/young/semistandard_tableau.lean @@ -8,6 +8,9 @@ import combinatorics.young.young_diagram /-! # Semistandard Young tableaux +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A semistandard Young tableau is a filling of a Young diagram by natural numbers, such that the entries are weakly increasing left-to-right along rows (i.e. for fixed `i`), and strictly-increasing top-to-bottom along columns (i.e. for fixed `j`). diff --git a/src/data/dfinsupp/order.lean b/src/data/dfinsupp/order.lean index 99b1df89079d9..0c6344b670dda 100644 --- a/src/data/dfinsupp/order.lean +++ b/src/data/dfinsupp/order.lean @@ -8,6 +8,9 @@ import data.dfinsupp.basic /-! # Pointwise order on finitely supported dependent functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file lifts order structures on the `α i` to `Π₀ i, α i`. ## Main declarations diff --git a/src/data/fin/tuple/bubble_sort_induction.lean b/src/data/fin/tuple/bubble_sort_induction.lean index b8e7a049b567c..8b496b9e8de09 100644 --- a/src/data/fin/tuple/bubble_sort_induction.lean +++ b/src/data/fin/tuple/bubble_sort_induction.lean @@ -10,6 +10,9 @@ import order.well_founded /-! # "Bubble sort" induction +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We implement the following induction principle `tuple.bubble_sort_induction` on tuples with values in a linear order `α`. diff --git a/src/data/multiset/fintype.lean b/src/data/multiset/fintype.lean index 83b15e8cb94ad..01ef958860e01 100644 --- a/src/data/multiset/fintype.lean +++ b/src/data/multiset/fintype.lean @@ -10,6 +10,9 @@ import data.prod.lex /-! # Multiset coercion to type +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This module defines a `has_coe_to_sort` instance for multisets and gives it a `fintype` instance. It also defines `multiset.to_enum_finset`, which is another way to enumerate the elements of a multiset. These coercions and definitions make it easier to sum over multisets using existing diff --git a/src/data/nat/part_enat.lean b/src/data/nat/part_enat.lean index b781908972ff0..7ce40ec7e0dea 100644 --- a/src/data/nat/part_enat.lean +++ b/src/data/nat/part_enat.lean @@ -11,6 +11,9 @@ import tactic.norm_num /-! # Natural numbers with infinity +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The natural numbers and an extra `top` element `⊤`. This implementation uses `part ℕ` as an implementation. Use `ℕ∞` instead unless you care about computability. diff --git a/src/data/pnat/factors.lean b/src/data/pnat/factors.lean index 10a72127cb164..805954aae0c34 100644 --- a/src/data/pnat/factors.lean +++ b/src/data/pnat/factors.lean @@ -12,6 +12,9 @@ import data.multiset.sort /-! # Prime factors of nonzero naturals +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the factorization of a nonzero natural number `n` as a multiset of primes, the multiplicity of `p` in this factors multiset being the p-adic valuation of `n`. diff --git a/src/data/setoid/partition.lean b/src/data/setoid/partition.lean index 187b22877bf4f..28919db4afddb 100644 --- a/src/data/setoid/partition.lean +++ b/src/data/setoid/partition.lean @@ -12,6 +12,9 @@ import order.partition.finpartition /-! # Equivalence relations: partitions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file comprises properties of equivalence relations viewed as partitions. There are two implementations of partitions here: * A collection `c : set (set α)` of sets is a partition of `α` if `∅ ∉ c` and each element `a : α` diff --git a/src/data/sigma/interval.lean b/src/data/sigma/interval.lean index 0c70e2cc3b00f..b28ad6a754d30 100644 --- a/src/data/sigma/interval.lean +++ b/src/data/sigma/interval.lean @@ -9,6 +9,9 @@ import order.locally_finite /-! # Finite intervals in a sigma type +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file provides the `locally_finite_order` instance for the disjoint sum of orders `Σ i, α i` and calculates the cardinality of its finite intervals. diff --git a/src/group_theory/double_coset.lean b/src/group_theory/double_coset.lean index 7d05e5d7f9e58..f7a35dcafbaee 100644 --- a/src/group_theory/double_coset.lean +++ b/src/group_theory/double_coset.lean @@ -13,6 +13,9 @@ import tactic.group /-! # Double cosets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines double cosets for two subgroups `H K` of a group `G` and the quotient of `G` by the double coset relation, i.e. `H \ G / K`. We also prove that `G` can be writen as a disjoint union of the double cosets and that if one of `H` or `K` is the trivial group (i.e. `⊥` ) then diff --git a/src/group_theory/perm/subgroup.lean b/src/group_theory/perm/subgroup.lean index cf8f21c6a40c3..77db74a606c95 100644 --- a/src/group_theory/perm/subgroup.lean +++ b/src/group_theory/perm/subgroup.lean @@ -9,6 +9,9 @@ import group_theory.subgroup.finite /-! # Lemmas about subgroups within the permutations (self-equivalences) of a type `α` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file provides extra lemmas about some `subgroup`s that exist within `equiv.perm α`. `group_theory.subgroup` depends on `group_theory.perm.basic`, so these need to be in a separate file. diff --git a/src/measure_theory/measurable_space_def.lean b/src/measure_theory/measurable_space_def.lean index a36f71f12a64f..642af17ecd600 100644 --- a/src/measure_theory/measurable_space_def.lean +++ b/src/measure_theory/measurable_space_def.lean @@ -10,6 +10,9 @@ import order.disjointed /-! # Measurable spaces and measurable functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines measurable spaces and measurable functions. A measurable space is a set equipped with a σ-algebra, a collection of diff --git a/src/order/heyting/hom.lean b/src/order/heyting/hom.lean index b63b779369d74..a947d251b340a 100644 --- a/src/order/heyting/hom.lean +++ b/src/order/heyting/hom.lean @@ -8,6 +8,9 @@ import order.hom.lattice /-! # Heyting algebra morphisms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A Heyting homomorphism between two Heyting algebras is a bounded lattice homomorphism that preserves Heyting implication. diff --git a/src/order/hom/complete_lattice.lean b/src/order/hom/complete_lattice.lean index 6f1a44ed63de6..354e1d03d6106 100644 --- a/src/order/hom/complete_lattice.lean +++ b/src/order/hom/complete_lattice.lean @@ -9,6 +9,9 @@ import order.hom.lattice /-! # Complete lattice homomorphisms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines frame homorphisms and complete lattice homomorphisms. We use the `fun_like` design, so each type of morphisms has a companion typeclass which is meant to diff --git a/src/topology/G_delta.lean b/src/topology/G_delta.lean index f9f07550a2592..d083170bcbafd 100644 --- a/src/topology/G_delta.lean +++ b/src/topology/G_delta.lean @@ -9,6 +9,9 @@ import topology.separation /-! # `Gδ` sets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define `Gδ` sets and prove their basic properties. ## Main definitions diff --git a/src/topology/algebra/order/archimedean.lean b/src/topology/algebra/order/archimedean.lean index bbf0859725ce2..33cdf711c7c15 100644 --- a/src/topology/algebra/order/archimedean.lean +++ b/src/topology/algebra/order/archimedean.lean @@ -9,6 +9,9 @@ import algebra.order.archimedean /-! # Rational numbers are dense in a linear ordered archimedean field +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove that coercion from `ℚ` to a linear ordered archimedean field has dense range. This lemma is in a separate file because `topology.order.basic` does not import `algebra.order.archimedean`. diff --git a/src/topology/algebra/order/compact.lean b/src/topology/algebra/order/compact.lean index 153b9fe5e7bf6..c7e46e18e122e 100644 --- a/src/topology/algebra/order/compact.lean +++ b/src/topology/algebra/order/compact.lean @@ -9,6 +9,9 @@ import topology.local_extr /-! # Compactness of a closed interval +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + 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. diff --git a/src/topology/algebra/order/extr_closure.lean b/src/topology/algebra/order/extr_closure.lean index 0d4cdddf5e2b8..89204e085b21e 100644 --- a/src/topology/algebra/order/extr_closure.lean +++ b/src/topology/algebra/order/extr_closure.lean @@ -9,6 +9,9 @@ import topology.order.basic /-! # Maximum/minimum on the closure of a set +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove several versions of the following statement: if `f : X → Y` has a (local or not) maximum (or minimum) on a set `s` at a point `a` and is continuous on the closure of `s`, then `f` has an extremum of the same type on `closure s` at `a`. diff --git a/src/topology/algebra/order/intermediate_value.lean b/src/topology/algebra/order/intermediate_value.lean index e82eba6907532..5e8cce657972d 100644 --- a/src/topology/algebra/order/intermediate_value.lean +++ b/src/topology/algebra/order/intermediate_value.lean @@ -9,6 +9,9 @@ import topology.order.basic /-! # Intermediate Value Theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove the Intermediate Value Theorem: if `f : α → β` is a function defined on a connected set `s` that takes both values `≤ a` and values `≥ a` on `s`, then it is equal to `a` at some point of `s`. We also prove that intervals in a dense conditionally complete order are diff --git a/src/topology/algebra/order/monotone_continuity.lean b/src/topology/algebra/order/monotone_continuity.lean index 1b6e22bd0e399..4671266286a4d 100644 --- a/src/topology/algebra/order/monotone_continuity.lean +++ b/src/topology/algebra/order/monotone_continuity.lean @@ -9,6 +9,9 @@ import topology.homeomorph /-! # Continuity of monotone functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove the following fact: if `f` is a monotone function on a neighborhood of `a` and the image of this neighborhood is a neighborhood of `f a`, then `f` is continuous at `a`, see `continuous_at_of_monotone_on_of_image_mem_nhds`, as well as several similar facts. diff --git a/src/topology/algebra/order/monotone_convergence.lean b/src/topology/algebra/order/monotone_convergence.lean index 10724d464f353..aae01923dbf9f 100644 --- a/src/topology/algebra/order/monotone_convergence.lean +++ b/src/topology/algebra/order/monotone_convergence.lean @@ -8,6 +8,9 @@ import topology.order.basic /-! # Bounded monotone sequences converge +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove a few theorems of the form “if the range of a monotone function `f : ι → α` admits a least upper bound `a`, then `f x` tends to `a` as `x → ∞`”, as well as version of this statement for (conditionally) complete lattices that use `⨆ x, f x` instead of `is_lub`. diff --git a/src/topology/algebra/order/proj_Icc.lean b/src/topology/algebra/order/proj_Icc.lean index c48e97502ec2a..0885475c3d961 100644 --- a/src/topology/algebra/order/proj_Icc.lean +++ b/src/topology/algebra/order/proj_Icc.lean @@ -9,6 +9,9 @@ import topology.order.basic /-! # Projection onto a closed interval +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove that the projection `set.proj_Icc f a b h` is a quotient map, and use it to show that `Icc_extend h f` is continuous if and only if `f` is continuous. -/ diff --git a/src/topology/algebra/order/t5.lean b/src/topology/algebra/order/t5.lean index 2e3d851be9fcf..14ac6d4ad9c25 100644 --- a/src/topology/algebra/order/t5.lean +++ b/src/topology/algebra/order/t5.lean @@ -9,6 +9,9 @@ import data.set.intervals.ord_connected_component /-! # Linear order is a completely normal Hausdorff topological space +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove that a linear order with order topology is a completely normal Hausdorff topological space. -/ diff --git a/src/topology/algebra/star.lean b/src/topology/algebra/star.lean index 99421351dbf6f..d05c702d09460 100644 --- a/src/topology/algebra/star.lean +++ b/src/topology/algebra/star.lean @@ -11,6 +11,9 @@ import topology.continuous_function.basic /-! # Continuity of `star` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the `has_continuous_star` typeclass, along with instances on `pi`, `prod`, `mul_opposite`, and `units`. -/ diff --git a/src/topology/continuous_function/basic.lean b/src/topology/continuous_function/basic.lean index 79d4178ec11fa..89c54869af349 100644 --- a/src/topology/continuous_function/basic.lean +++ b/src/topology/continuous_function/basic.lean @@ -10,6 +10,9 @@ import topology.homeomorph /-! # Continuous bundled maps +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define the type `continuous_map` of continuous bundled maps. We use the `fun_like` design, so each type of morphisms has a companion typeclass which is meant to diff --git a/src/topology/instances/sign.lean b/src/topology/instances/sign.lean index bcd69e311abe4..2ddc6e4e4d487 100644 --- a/src/topology/instances/sign.lean +++ b/src/topology/instances/sign.lean @@ -9,6 +9,9 @@ import topology.order.basic /-! # Topology on `sign_type` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file gives `sign_type` the discrete topology, and proves continuity results for `sign` in an `order_topology`. diff --git a/src/topology/order/priestley.lean b/src/topology/order/priestley.lean index 15355ae5006c2..8cbdc51758425 100644 --- a/src/topology/order/priestley.lean +++ b/src/topology/order/priestley.lean @@ -9,6 +9,9 @@ import topology.separation /-! # Priestley spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines Priestley spaces. A Priestley space is an ordered compact topological space such that any two distinct points can be separated by a clopen upper set. diff --git a/src/topology/perfect.lean b/src/topology/perfect.lean index 051687e4c4dba..24609fd926296 100644 --- a/src/topology/perfect.lean +++ b/src/topology/perfect.lean @@ -9,6 +9,9 @@ import topology.bases /-! # Perfect Sets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define perfect subsets of a topological space, and prove some basic properties, including a version of the Cantor-Bendixson Theorem. From 0eb49606a00a1eb20d7238a32169dca2a660f261 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Tue, 7 Feb 2023 07:32:09 +0000 Subject: [PATCH 0433/1291] chore(topology/noetherian_space): golf (#18394) Also generalize `noetherian_space.set` to `inducing.noetherian_space`. --- src/topology/noetherian_space.lean | 36 +++++++++++++----------------- 1 file changed, 15 insertions(+), 21 deletions(-) diff --git a/src/topology/noetherian_space.lean b/src/topology/noetherian_space.lean index 1d2ad9f3991d0..2a1f3fda92787 100644 --- a/src/topology/noetherian_space.lean +++ b/src/topology/noetherian_space.lean @@ -54,28 +54,25 @@ end @[priority 100] instance noetherian_space.compact_space [h : noetherian_space α] : compact_space α := -is_compact_univ_iff.mp ((noetherian_space_iff_opens α).mp h ⊤) +⟨(noetherian_space_iff_opens α).mp h ⊤⟩ -variable {α} +variables {α β} -instance noetherian_space.set [h : noetherian_space α] (s : set α) : noetherian_space s := +protected lemma noetherian_space.is_compact [noetherian_space α] (s : set α) : is_compact s := begin - rw noetherian_space_iff, - apply well_founded.well_founded_iff_has_max'.2, - intros p hp, - obtain ⟨⟨_, u, hu, rfl⟩, hu'⟩ := hp, - obtain ⟨U, hU, hU'⟩ := well_founded.well_founded_iff_has_max'.1 h.1 - (((opens.comap ⟨_, continuous_subtype_coe⟩)) ⁻¹' p) ⟨⟨u, hu⟩, hu'⟩, - refine ⟨opens.comap ⟨_, continuous_subtype_coe⟩ U, hU, _⟩, - rintros ⟨_, x, hx, rfl⟩ hx' hx'', - refine le_antisymm (set.preimage_mono (_ : (⟨x, hx⟩ : opens α) ≤ U)) hx'', - refine sup_eq_right.mp (hU' (⟨x, hx⟩ ⊔ U) _ le_sup_right), - dsimp [set.preimage], - rw map_sup, - convert hx', - exact sup_eq_left.mpr hx'' + refine is_compact_iff_finite_subcover.2 (λ ι U hUo hs, _), + rcases ((noetherian_space_iff_opens α).mp ‹_› + ⟨⋃ i, U i, is_open_Union hUo⟩).elim_finite_subcover U hUo set.subset.rfl with ⟨t, ht⟩, + exact ⟨t, hs.trans ht⟩ end +protected lemma inducing.noetherian_space [noetherian_space α] {i : β → α} (hi : inducing i) : + noetherian_space β := +(noetherian_space_iff_opens _).2 $ λ s, hi.is_compact_iff.1 (noetherian_space.is_compact _) + +instance noetherian_space.set [h : noetherian_space α] (s : set α) : noetherian_space s := +inducing_coe.noetherian_space + variable (α) example (α : Type*) : set α ≃o (set α)ᵒᵈ := by refine order_iso.compl (set α) @@ -92,7 +89,7 @@ begin tfae_have : 1 ↔ 4, { exact noetherian_space_iff_opens α }, tfae_have : 1 → 3, - { intros H s, rw is_compact_iff_compact_space, resetI, apply_instance }, + { exact @noetherian_space.is_compact _ _ }, tfae_have : 3 → 4, { exact λ H s, H s }, tfae_finish @@ -111,9 +108,6 @@ begin { exact ⟨a, filter.le_principal_iff.mp hs, or.inl le_rfl⟩ } end -lemma noetherian_space.is_compact [h : noetherian_space α] (s : set α) : is_compact s := -let H := (noetherian_space_tfae α).out 0 2 in H.mp h s - lemma noetherian_space_of_surjective [noetherian_space α] (f : α → β) (hf : continuous f) (hf' : function.surjective f) : noetherian_space β := begin From 50092e4a9fda259fe1c6430468f3fac003d43f56 Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Tue, 7 Feb 2023 09:01:49 +0000 Subject: [PATCH 0434/1291] feat(analysis/convolution): integral of a convolution over positive reals (#18353) This PR generalises `integral_convolution` to consider convolutions of functions on the positive real line. (This will be used in a follow-up PR to relate the gamma and beta functions). --- src/analysis/convolution.lean | 73 +++++++++++++++++++++++++++++++++++ 1 file changed, 73 insertions(+) diff --git a/src/analysis/convolution.lean b/src/analysis/convolution.lean index 64b87646fa83f..f21b490b15f2e 100644 --- a/src/analysis/convolution.lean +++ b/src/analysis/convolution.lean @@ -6,6 +6,7 @@ Authors: Floris van Doorn import measure_theory.group.integration import measure_theory.group.prod import measure_theory.function.locally_integrable +import measure_theory.integral.interval_integral import analysis.calculus.specific_functions import analysis.calculus.parametric_integral @@ -1562,3 +1563,75 @@ lemma has_compact_support.cont_diff_convolution_left [μ.is_add_left_invariant] by { rw [← convolution_flip], exact hcf.cont_diff_convolution_right L.flip hg hf } end with_param + +section nonneg + +variables [normed_space ℝ E] [normed_space ℝ E'] [normed_space ℝ F] [complete_space F] + +/-- The forward convolution of two functions `f` and `g` on `ℝ`, with respect to a continuous +bilinear map `L` and measure `ν`. It is defined to be the function mapping `x` to +`∫ t in 0..x, L (f t) (g (x - t)) ∂ν` if `0 < x`, and 0 otherwise. -/ +noncomputable def pos_convolution + (f : ℝ → E) (g : ℝ → E') (L : E →L[ℝ] E' →L[ℝ] F) (ν : measure ℝ . volume_tac) : ℝ → F := +indicator (Ioi (0:ℝ)) (λ x, ∫ t in 0..x, L (f t) (g (x - t)) ∂ν) + +lemma pos_convolution_eq_convolution_indicator + (f : ℝ → E) (g : ℝ → E') (L : E →L[ℝ] E' →L[ℝ] F) (ν : measure ℝ . volume_tac) [has_no_atoms ν] : + pos_convolution f g L ν = convolution (indicator (Ioi 0) f) (indicator (Ioi 0) g) L ν := +begin + ext1 x, + rw [convolution, pos_convolution, indicator], + split_ifs, + { rw [interval_integral.integral_of_le (le_of_lt h), + integral_Ioc_eq_integral_Ioo, + ←integral_indicator (measurable_set_Ioo : measurable_set (Ioo 0 x))], + congr' 1 with t : 1, + have : (t ≤ 0) ∨ (t ∈ Ioo 0 x) ∨ (x ≤ t), + { rcases le_or_lt t 0 with h | h, + { exact or.inl h }, + { rcases lt_or_le t x with h' | h', + exacts [or.inr (or.inl ⟨h, h'⟩), or.inr (or.inr h')] } }, + rcases this with ht|ht|ht, + { rw [indicator_of_not_mem (not_mem_Ioo_of_le ht), indicator_of_not_mem (not_mem_Ioi.mpr ht), + continuous_linear_map.map_zero, continuous_linear_map.zero_apply] }, + { rw [indicator_of_mem ht, indicator_of_mem (mem_Ioi.mpr ht.1), + indicator_of_mem (mem_Ioi.mpr $ sub_pos.mpr ht.2)] }, + { rw [indicator_of_not_mem (not_mem_Ioo_of_ge ht), + indicator_of_not_mem (not_mem_Ioi.mpr (sub_nonpos_of_le ht)), + continuous_linear_map.map_zero] } }, + { convert (integral_zero ℝ F).symm, + ext1 t, + by_cases ht : 0 < t, + { rw [indicator_of_not_mem (_ : x - t ∉ Ioi 0), continuous_linear_map.map_zero], + rw not_mem_Ioi at h ⊢, + exact sub_nonpos.mpr (h.trans ht.le) }, + { rw [indicator_of_not_mem (mem_Ioi.not.mpr ht), continuous_linear_map.map_zero, + continuous_linear_map.zero_apply] } } +end + +lemma integrable_pos_convolution {f : ℝ → E} {g : ℝ → E'} {μ ν : measure ℝ} + [sigma_finite μ] [sigma_finite ν] [is_add_right_invariant μ] [has_no_atoms ν] + (hf : integrable_on f (Ioi 0) ν) (hg : integrable_on g (Ioi 0) μ) (L : E →L[ℝ] E' →L[ℝ] F) : + integrable (pos_convolution f g L ν) μ := +begin + rw ←integrable_indicator_iff (measurable_set_Ioi : measurable_set (Ioi (0:ℝ))) at hf hg, + rw pos_convolution_eq_convolution_indicator f g L ν, + exact (hf.convolution_integrand L hg).integral_prod_left, +end + +/-- The integral over `Ioi 0` of a forward convolution of two functions is equal to the product +of their integrals over this set. (Compare `integral_convolution` for the two-sided convolution.) -/ +lemma integral_pos_convolution [complete_space E] [complete_space E'] {μ ν : measure ℝ} + [sigma_finite μ] [sigma_finite ν] [is_add_right_invariant μ] [has_no_atoms ν] + {f : ℝ → E} {g : ℝ → E'} (hf : integrable_on f (Ioi 0) ν) + (hg : integrable_on g (Ioi 0) μ) (L : E →L[ℝ] E' →L[ℝ] F) : + ∫ x:ℝ in Ioi 0, (∫ t:ℝ in 0..x, L (f t) (g (x - t)) ∂ν) ∂μ = + L (∫ x:ℝ in Ioi 0, f x ∂ν) (∫ x:ℝ in Ioi 0, g x ∂μ) := +begin + rw ←integrable_indicator_iff (measurable_set_Ioi : measurable_set (Ioi (0:ℝ))) at hf hg, + simp_rw ←integral_indicator measurable_set_Ioi, + convert integral_convolution L hf hg using 2, + apply pos_convolution_eq_convolution_indicator, +end + +end nonneg From 98e83c3d541c77cdb7da20d79611a780ff8e7d90 Mon Sep 17 00:00:00 2001 From: negiizhao Date: Tue, 7 Feb 2023 20:26:31 +0000 Subject: [PATCH 0435/1291] feat(set_theory/zfc/basic): define `hereditarily` (#18328) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We will define von Neumann ordinals as hereditarily transitive sets in #18329. Also there are [hereditarily finite sets](https://en.wikipedia.org/wiki/Hereditarily_finite_set) and [hereditarily countable sets](https://en.wikipedia.org/wiki/Hereditarily_countable_set) in set theory. Co-authored-by: Violeta Hernández [vi.hdz.p@gmail.com](mailto:vi.hdz.p@gmail.com) --- src/set_theory/zfc/basic.lean | 32 ++++++++++++++++++++++++++++++++ 1 file changed, 32 insertions(+) diff --git a/src/set_theory/zfc/basic.lean b/src/set_theory/zfc/basic.lean index d5b8d9422ab14..7396b9188f529 100644 --- a/src/set_theory/zfc/basic.lean +++ b/src/set_theory/zfc/basic.lean @@ -44,6 +44,8 @@ Then the rest is usual set theory. function `x → y`. That is, each member of `x` is related by the ZFC set to exactly one member of `y`. * `Set.funs`: ZFC set of ZFC functions `x → y`. +* `Set.hereditarily p x`: Predicate that every set in the transitive closure of `x` has property + `p`. * `Class.iota`: Definite description operator. ## Notes @@ -857,6 +859,36 @@ theorem map_unique {f : Set.{u} → Set.{u}} [H : definable 1 f] {x z : Set.{u}} λ h, ⟨λ y yx, let ⟨z, zx, ze⟩ := mem_image.1 yx in ze ▸ pair_mem_prod.2 ⟨zx, h z zx⟩, λ z, map_unique⟩⟩ +/-- Given a predicate `p` on ZFC sets. `hereditarily p x` means that `x` has property `p` and the +members of `x` are all `hereditarily p`. -/ +def hereditarily (p : Set → Prop) : Set → Prop +| x := p x ∧ ∀ y ∈ x, hereditarily y +using_well_founded { dec_tac := `[assumption] } + +section hereditarily + +variables {p : Set.{u} → Prop} {x y : Set.{u}} + +lemma hereditarily_iff : + hereditarily p x ↔ p x ∧ ∀ y ∈ x, hereditarily p y := +by rw [← hereditarily] + +alias hereditarily_iff ↔ hereditarily.def _ + +lemma hereditarily.self (h : x.hereditarily p) : p x := h.def.1 +lemma hereditarily.mem (h : x.hereditarily p) (hy : y ∈ x) : y.hereditarily p := h.def.2 _ hy + +lemma hereditarily.empty : hereditarily p x → p ∅ := +begin + apply x.induction_on, + intros y IH h, + rcases Set.eq_empty_or_nonempty y with (rfl|⟨a, ha⟩), + { exact h.self }, + { exact IH a ha (h.mem ha) } +end + +end hereditarily + end Set /-- The collection of all classes. From 3e32bc908f617039c74c06ea9a897e30c30803c2 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Wed, 8 Feb 2023 07:36:01 +0000 Subject: [PATCH 0436/1291] chore(*): add mathlib4 synchronization comments (#18398) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.add_torsor` * `category_theory.bicategory.basic` * `data.finsupp.antidiagonal` * `data.nat.choose.central` * `data.nat.choose.sum` * `number_theory.primorial` * `order.liminf_limsup` * `topology.algebra.order.extend_from` * `topology.continuous_function.cocompact_map` * `topology.locally_constant.basic` * `topology.order.lattice` * `topology.sober` * `topology.stone_cech` --- src/algebra/add_torsor.lean | 3 +++ src/category_theory/bicategory/basic.lean | 3 +++ src/data/finsupp/antidiagonal.lean | 3 +++ src/data/nat/choose/central.lean | 3 +++ src/data/nat/choose/sum.lean | 3 +++ src/number_theory/primorial.lean | 3 +++ src/order/liminf_limsup.lean | 3 +++ src/topology/algebra/order/extend_from.lean | 3 +++ src/topology/continuous_function/cocompact_map.lean | 3 +++ src/topology/locally_constant/basic.lean | 3 +++ src/topology/order/lattice.lean | 3 +++ src/topology/sober.lean | 3 +++ src/topology/stone_cech.lean | 3 +++ 13 files changed, 39 insertions(+) diff --git a/src/algebra/add_torsor.lean b/src/algebra/add_torsor.lean index 09deb0a0fe63f..6b72f08f2be91 100644 --- a/src/algebra/add_torsor.lean +++ b/src/algebra/add_torsor.lean @@ -8,6 +8,9 @@ import data.set.pointwise.smul /-! # Torsors of additive group actions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines torsors of additive group actions. ## Notations diff --git a/src/category_theory/bicategory/basic.lean b/src/category_theory/bicategory/basic.lean index c69c681a711bf..890c3299d03af 100644 --- a/src/category_theory/bicategory/basic.lean +++ b/src/category_theory/bicategory/basic.lean @@ -9,6 +9,9 @@ import tactic.slice /-! # Bicategories +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define typeclass for bicategories. A bicategory `B` consists of diff --git a/src/data/finsupp/antidiagonal.lean b/src/data/finsupp/antidiagonal.lean index 769145608c0b8..388c8dcdfb787 100644 --- a/src/data/finsupp/antidiagonal.lean +++ b/src/data/finsupp/antidiagonal.lean @@ -9,6 +9,9 @@ import data.multiset.antidiagonal /-! # The `finsupp` counterpart of `multiset.antidiagonal`. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The antidiagonal of `s : α →₀ ℕ` consists of all pairs `(t₁, t₂) : (α →₀ ℕ) × (α →₀ ℕ)` such that `t₁ + t₂ = s`. -/ diff --git a/src/data/nat/choose/central.lean b/src/data/nat/choose/central.lean index 10b2f90d9849f..c6755734d8394 100644 --- a/src/data/nat/choose/central.lean +++ b/src/data/nat/choose/central.lean @@ -10,6 +10,9 @@ import tactic.linarith /-! # Central binomial coefficients +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves properties of the central binomial coefficients (that is, `nat.choose (2 * n) n`). ## Main definition and results diff --git a/src/data/nat/choose/sum.lean b/src/data/nat/choose/sum.lean index fb9d4b7488df9..2304080d66257 100644 --- a/src/data/nat/choose/sum.lean +++ b/src/data/nat/choose/sum.lean @@ -13,6 +13,9 @@ import algebra.big_operators.nat_antidiagonal /-! # Sums of binomial coefficients +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file includes variants of the binomial theorem and other results on sums of binomial coefficients. Theorems whose proofs depend on such sums may also go in this file for import reasons. diff --git a/src/number_theory/primorial.lean b/src/number_theory/primorial.lean index f216faaed13a3..5eea76e716d1e 100644 --- a/src/number_theory/primorial.lean +++ b/src/number_theory/primorial.lean @@ -12,6 +12,9 @@ import data.nat.prime /-! # Primorial +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the primorial function (the product of primes less than or equal to some bound), and proves that `primorial n ≤ 4 ^ n`. diff --git a/src/order/liminf_limsup.lean b/src/order/liminf_limsup.lean index 8540b8d61e19f..adafe31cd2008 100644 --- a/src/order/liminf_limsup.lean +++ b/src/order/liminf_limsup.lean @@ -9,6 +9,9 @@ import order.hom.complete_lattice /-! # liminfs and limsups of functions and filters +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Defines the Liminf/Limsup of a function taking values in a conditionally complete lattice, with respect to an arbitrary filter. diff --git a/src/topology/algebra/order/extend_from.lean b/src/topology/algebra/order/extend_from.lean index 41517fd80e2b5..d18d2ddd35355 100644 --- a/src/topology/algebra/order/extend_from.lean +++ b/src/topology/algebra/order/extend_from.lean @@ -8,6 +8,9 @@ import topology.extend_from /-! # Lemmas about `extend_from` in an order topology. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ open filter set topological_space diff --git a/src/topology/continuous_function/cocompact_map.lean b/src/topology/continuous_function/cocompact_map.lean index 53042c5114a23..7b5b64c25d702 100644 --- a/src/topology/continuous_function/cocompact_map.lean +++ b/src/topology/continuous_function/cocompact_map.lean @@ -8,6 +8,9 @@ import topology.continuous_function.basic /-! # Cocompact continuous maps +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The type of *cocompact continuous maps* are those which tend to the cocompact filter on the codomain along the cocompact filter on the domain. When the domain and codomain are Hausdorff, this is equivalent to many other conditions, including that preimages of compact sets are compact. -/ diff --git a/src/topology/locally_constant/basic.lean b/src/topology/locally_constant/basic.lean index 436f30749555f..3a46a0fce005e 100644 --- a/src/topology/locally_constant/basic.lean +++ b/src/topology/locally_constant/basic.lean @@ -13,6 +13,9 @@ import tactic.fin_cases /-! # Locally constant functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file sets up the theory of locally constant function from a topological space to a type. ## Main definitions and constructions diff --git a/src/topology/order/lattice.lean b/src/topology/order/lattice.lean index 6817698750018..4c189d437a7a5 100644 --- a/src/topology/order/lattice.lean +++ b/src/topology/order/lattice.lean @@ -9,6 +9,9 @@ import topology.constructions /-! # Topological lattices +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define mixin classes `has_continuous_inf` and `has_continuous_sup`. We define the class `topological_lattice` as a topological space and lattice `L` extending `has_continuous_inf` and `has_continuous_sup`. diff --git a/src/topology/sober.lean b/src/topology/sober.lean index 58252d9a8ef06..853321809143a 100644 --- a/src/topology/sober.lean +++ b/src/topology/sober.lean @@ -8,6 +8,9 @@ import topology.separation /-! # Sober spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A quasi-sober space is a topological space where every irreducible closed subset has a generic point. A sober space is a quasi-sober space where every irreducible closed subset diff --git a/src/topology/stone_cech.lean b/src/topology/stone_cech.lean index b5410a75a9467..d8f7cc15fb3bc 100644 --- a/src/topology/stone_cech.lean +++ b/src/topology/stone_cech.lean @@ -8,6 +8,9 @@ import topology.dense_embedding /-! # Stone-Čech compactification +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Construction of the Stone-Čech compactification using ultrafilters. Parts of the formalization are based on "Ultrafilters and Topology" From b018406ad2f2a73223a3a9e198ccae61e6f05318 Mon Sep 17 00:00:00 2001 From: sgouezel Date: Wed, 8 Feb 2023 10:26:23 +0000 Subject: [PATCH 0437/1291] feat(analysis/calculus/special_functions): refactor bump functions (#17453) Currently, we have satisfactory bump functions on inner spaces (equal to 1 on `ball c r` and support equal to `ball c R`), and less satisfactory bump functions on general finite-dimensional normed spaces (equal to 1 on a neighborhood of `c`, with support a larger neighborhood of `c`), where the latter are obtained from the former by composing with an arbitrary linear equiv with an inner product space called `to_euclidean`. In particular, they have different names (`bump_function_of_inner` and `bump_function`) and whenever one wants to prove properties of `bump_function` one has to juggle with the `to_euclidean`. We refactor bump functions to get a coherent theory, which is uniform over all normed spaces. We remove completely `bump_function_of_inner`, and we make sure that `bump_function c r R` is a smooth function equal to `1` on `ball c r` and with support exactly `ball c R`, on all normed spaces. For this, we provide a typeclass `has_cont_diff_bump` saying that a space has nice bump functions, and we instantiate it both on inner product spaces (using the old approach with a function constructed from the norm and from `exp(-1/x^2)`), and on finite-dimensional normed spaces (where the bump functions have already been constructed by convolution in #18095). --- docs/overview.yaml | 2 +- docs/undergrad.yaml | 2 +- .../calculus/bump_function_findim.lean | 28 +- ...unctions.lean => bump_function_inner.lean} | 306 +++++++++--------- src/analysis/convolution.lean | 20 +- src/geometry/manifold/bump_function.lean | 43 ++- 6 files changed, 181 insertions(+), 220 deletions(-) rename src/analysis/calculus/{specific_functions.lean => bump_function_inner.lean} (67%) diff --git a/docs/overview.yaml b/docs/overview.yaml index 4c85da0604c81..1857a175084bf 100644 --- a/docs/overview.yaml +++ b/docs/overview.yaml @@ -337,7 +337,7 @@ Analysis: Fubini's theorem: 'measure_theory.integral_prod' product of finitely many measures: 'measure_theory.measure.pi' convolution: 'convolution' - approximation by convolution: 'cont_diff_bump_of_inner.convolution_tendsto_right' + approximation by convolution: 'cont_diff_bump.convolution_tendsto_right' regularization by convolution: 'has_compact_support.cont_diff_convolution_left' change of variables formula: 'measure_theory.integral_image_eq_integral_abs_det_fderiv_smul' divergence theorem: 'measure_theory.integral_divergence_of_has_fderiv_within_at_off_countable' diff --git a/docs/undergrad.yaml b/docs/undergrad.yaml index 18d92c12e930f..c984b13916c83 100644 --- a/docs/undergrad.yaml +++ b/docs/undergrad.yaml @@ -518,7 +518,7 @@ Measures and integral calculus: change of variables to polar co-ordinates: 'integral_comp_polar_coord_symm ' change of variables to spherical co-ordinates: '' convolution: 'convolution' - approximation by convolution: 'cont_diff_bump_of_inner.convolution_tendsto_right' + approximation by convolution: 'cont_diff_bump.convolution_tendsto_right' regularization by convolution: 'has_compact_support.cont_diff_convolution_left' Fourier analysis: Fourier series of locally integrable periodic real-valued functions: 'fourier_basis' diff --git a/src/analysis/calculus/bump_function_findim.lean b/src/analysis/calculus/bump_function_findim.lean index 0ea049edb8028..848daeca1936b 100644 --- a/src/analysis/calculus/bump_function_findim.lean +++ b/src/analysis/calculus/bump_function_findim.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ -import analysis.calculus.specific_functions +import analysis.calculus.bump_function_inner import analysis.calculus.series import analysis.convolution import data.set.pointwise.support @@ -16,9 +16,8 @@ Let `E` be a finite-dimensional real normed vector space. We show that any open exactly the support of a smooth function taking values in `[0, 1]`, in `is_open.exists_smooth_support_eq`. -TODO: use this construction to construct bump functions with nice behavior in any finite-dimensional -real normed vector space, by convolving the indicator function of `closed_ball 0 1` with a -function as above with `s = ball 0 D`. +Then we use this construction to construct bump functions with nice behavior, by convolving +the indicator function of `closed_ball 0 1` with a function as above with `s = ball 0 D`. -/ noncomputable theory @@ -41,7 +40,7 @@ theorem exists_smooth_tsupport_subset {s : set E} {x : E} (hs : s ∈ 𝓝 x) : begin obtain ⟨d, d_pos, hd⟩ : ∃ (d : ℝ) (hr : 0 < d), euclidean.closed_ball x d ⊆ s, from euclidean.nhds_basis_closed_ball.mem_iff.1 hs, - let c : cont_diff_bump_of_inner (to_euclidean x) := + let c : cont_diff_bump (to_euclidean x) := { r := d/2, R := d, r_pos := half_pos d_pos, @@ -478,25 +477,6 @@ variable {E} end helper_definitions -/-- The base function from which one will construct a family of bump functions. One could -add more properties if they are useful and satisfied in the examples of inner product spaces -and finite dimensional vector spaces, notably derivative norm control in terms of `R - 1`. -TODO: move to the right file and use this to refactor bump functions in general. -/ -@[nolint has_nonempty_instance] -structure _root_.cont_diff_bump_base (E : Type*) [normed_add_comm_group E] [normed_space ℝ E] := -(to_fun : ℝ → E → ℝ) -(mem_Icc : ∀ (R : ℝ) (x : E), to_fun R x ∈ Icc (0 : ℝ) 1) -(symmetric : ∀ (R : ℝ) (x : E), to_fun R (-x) = to_fun R x) -(smooth : cont_diff_on ℝ ⊤ (uncurry to_fun) ((Ioi (1 : ℝ)) ×ˢ (univ : set E))) -(eq_one : ∀ (R : ℝ) (hR : 1 < R) (x : E) (hx : ‖x‖ ≤ 1), to_fun R x = 1) -(support : ∀ (R : ℝ) (hR : 1 < R), support (to_fun R) = metric.ball (0 : E) R) - -/-- A class registering that a real vector space admits bump functions. This will be instantiated -first for inner product spaces, and then for finite-dimensional normed spaces. -We use a specific class instead of `nonempty (cont_diff_bump_base E)` for performance reasons. -/ -class _root_.has_cont_diff_bump (E : Type*) [normed_add_comm_group E] [normed_space ℝ E] : Prop := -(out : nonempty (cont_diff_bump_base E)) - @[priority 100] instance {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] [finite_dimensional ℝ E] : has_cont_diff_bump E := diff --git a/src/analysis/calculus/specific_functions.lean b/src/analysis/calculus/bump_function_inner.lean similarity index 67% rename from src/analysis/calculus/specific_functions.lean rename to src/analysis/calculus/bump_function_inner.lean index 7cca302310603..8a8987b7af8cd 100644 --- a/src/analysis/calculus/specific_functions.lean +++ b/src/analysis/calculus/bump_function_inner.lean @@ -20,29 +20,19 @@ function cannot have: * `real.smooth_transition` is equal to zero for `x ≤ 0` and is equal to one for `x ≥ 1`; it is given by `exp_neg_inv_glue x / (exp_neg_inv_glue x + exp_neg_inv_glue (1 - x))`; -* `f : cont_diff_bump_of_inner c`, where `c` is a point in an inner product space, is +* `f : cont_diff_bump c`, where `c` is a point in a real vector space, is a bundled smooth function such that - `f` is equal to `1` in `metric.closed_ball c f.r`; - `support f = metric.ball c f.R`; - `0 ≤ f x ≤ 1` for all `x`. - The structure `cont_diff_bump_of_inner` contains the data required to construct the + The structure `cont_diff_bump` contains the data required to construct the function: real numbers `r`, `R`, and proofs of `0 < r < R`. The function itself is available through `coe_fn`. -* If `f : cont_diff_bump_of_inner c` and `μ` is a measure on the domain of `f`, then `f.normed μ` +* If `f : cont_diff_bump c` and `μ` is a measure on the domain of `f`, then `f.normed μ` is a smooth bump function with integral `1` w.r.t. `μ`. - -* `f : cont_diff_bump c`, where `c` is a point in a finite dimensional real vector space, is a - bundled smooth function such that - - - `f` is equal to `1` in `euclidean.closed_ball c f.r`; - - `support f = euclidean.ball c f.R`; - - `0 ≤ f x ≤ 1` for all `x`. - - The structure `cont_diff_bump` contains the data required to construct the function: real - numbers `r`, `R`, and proofs of `0 < r < R`. The function itself is available through `coe_fn`. -/ noncomputable theory @@ -275,76 +265,173 @@ end real variables {E X : Type*} -/-- `f : cont_diff_bump_of_inner c`, where `c` is a point in an inner product space, is a +/-- `f : cont_diff_bump c`, where `c` is a point in a normed vector space, is a bundled smooth function such that - `f` is equal to `1` in `metric.closed_ball c f.r`; - `support f = metric.ball c f.R`; - `0 ≤ f x ≤ 1` for all `x`. -The structure `cont_diff_bump_of_inner` contains the data required to construct the function: +The structure `cont_diff_bump` contains the data required to construct the function: real numbers `r`, `R`, and proofs of `0 < r < R`. The function itself is available through -`coe_fn`. -/ -structure cont_diff_bump_of_inner (c : E) := +`coe_fn` when the space is nice enough, i.e., satisfies the `has_cont_diff_bump` typeclass. -/ +structure cont_diff_bump (c : E) := (r R : ℝ) (r_pos : 0 < r) (r_lt_R : r < R) -namespace cont_diff_bump_of_inner +/-- The base function from which one will construct a family of bump functions. One could +add more properties if they are useful and satisfied in the examples of inner product spaces +and finite dimensional vector spaces, notably derivative norm control in terms of `R - 1`. -/ +@[nolint has_nonempty_instance] +structure cont_diff_bump_base (E : Type*) [normed_add_comm_group E] [normed_space ℝ E] := +(to_fun : ℝ → E → ℝ) +(mem_Icc : ∀ (R : ℝ) (x : E), to_fun R x ∈ Icc (0 : ℝ) 1) +(symmetric : ∀ (R : ℝ) (x : E), to_fun R (-x) = to_fun R x) +(smooth : cont_diff_on ℝ ⊤ (uncurry to_fun) ((Ioi (1 : ℝ)) ×ˢ (univ : set E))) +(eq_one : ∀ (R : ℝ) (hR : 1 < R) (x : E) (hx : ‖x‖ ≤ 1), to_fun R x = 1) +(support : ∀ (R : ℝ) (hR : 1 < R), support (to_fun R) = metric.ball (0 : E) R) + +/-- A class registering that a real vector space admits bump functions. This will be instantiated +first for inner product spaces, and then for finite-dimensional normed spaces. +We use a specific class instead of `nonempty (cont_diff_bump_base E)` for performance reasons. -/ +class has_cont_diff_bump (E : Type*) [normed_add_comm_group E] [normed_space ℝ E] : Prop := +(out : nonempty (cont_diff_bump_base E)) + +/-- In a space with `C^∞` bump functions, register some function that will be used as a basis +to construct bump functions of arbitrary size around any point. -/ +def some_cont_diff_bump_base (E : Type*) [normed_add_comm_group E] [normed_space ℝ E] + [hb : has_cont_diff_bump E] : cont_diff_bump_base E := +nonempty.some hb.out + +/-- Any inner product space has smooth bump functions. -/ +@[priority 100] instance has_cont_diff_bump_of_inner_product_space + (E : Type*) [inner_product_space ℝ E] : has_cont_diff_bump E := +let e : cont_diff_bump_base E := +{ to_fun := λ R x, real.smooth_transition ((R - ‖x‖) / (R - 1)), + mem_Icc := λ R x, ⟨real.smooth_transition.nonneg _, real.smooth_transition.le_one _⟩, + symmetric := λ R x, by simp only [norm_neg], + smooth := begin + rintros ⟨R, x⟩ ⟨(hR : 1 < R), hx⟩, + apply cont_diff_at.cont_diff_within_at, + rcases eq_or_ne x 0 with rfl|hx, + { have : (λ (p : ℝ × E), real.smooth_transition ((p.1 - ‖p.2‖) / (p.1 - 1))) + =ᶠ[𝓝 (R, 0)] (λ p, 1), + { have A : tendsto (λ (p : ℝ × E), (p.1 - ‖p.2‖) / (p.1 - 1)) + (𝓝 (R, 0)) (𝓝 ((R - ‖(0 : E)‖) / (R - 1))), + { rw nhds_prod_eq, + apply (tendsto_fst.sub tendsto_snd.norm).div (tendsto_fst.sub tendsto_const_nhds), + exact (sub_pos.2 hR).ne' }, + have : ∀ᶠ (p : ℝ × E) in 𝓝 (R, 0), 1 < (p.1 - ‖p.2‖) / (p.1 - 1), + { apply (tendsto_order.1 A).1, + apply (one_lt_div (sub_pos.2 hR)).2, + simp only [norm_zero, tsub_zero, sub_lt_self_iff, zero_lt_one] }, + filter_upwards [this] with q hq, + exact real.smooth_transition.one_of_one_le hq.le }, + exact cont_diff_at_const.congr_of_eventually_eq this }, + { refine real.smooth_transition.cont_diff_at.comp _ _, + refine cont_diff_at.div _ _ (sub_pos.2 hR).ne', + { exact cont_diff_at_fst.sub (cont_diff_at_snd.norm hx) }, + { exact cont_diff_at_fst.sub cont_diff_at_const } } + end, + eq_one := λ R hR x hx, real.smooth_transition.one_of_one_le $ + (one_le_div (sub_pos.2 hR)).2 (sub_le_sub_left hx _), + support := λ R hR, begin + apply subset.antisymm, + { assume x hx, + simp only [mem_support] at hx, + contrapose! hx, + simp only [mem_ball_zero_iff, not_lt] at hx, + apply real.smooth_transition.zero_of_nonpos, + apply div_nonpos_of_nonpos_of_nonneg; + linarith }, + { assume x hx, + simp only [mem_ball_zero_iff] at hx, + apply (real.smooth_transition.pos_of_pos _).ne', + apply div_pos; + linarith } + end, } +in ⟨⟨e⟩⟩ + +namespace cont_diff_bump + +lemma R_pos {c : E} (f : cont_diff_bump c) : 0 < f.R := f.r_pos.trans f.r_lt_R -lemma R_pos {c : E} (f : cont_diff_bump_of_inner c) : 0 < f.R := f.r_pos.trans f.r_lt_R +lemma one_lt_R_div_r {c : E} (f : cont_diff_bump c) : 1 < f.R / f.r := +begin + rw one_lt_div f.r_pos, + exact f.r_lt_R +end -instance (c : E) : inhabited (cont_diff_bump_of_inner c) := ⟨⟨1, 2, zero_lt_one, one_lt_two⟩⟩ +instance (c : E) : inhabited (cont_diff_bump c) := ⟨⟨1, 2, zero_lt_one, one_lt_two⟩⟩ -variables [inner_product_space ℝ E] [normed_add_comm_group X] [normed_space ℝ X] -variables {c : E} (f : cont_diff_bump_of_inner c) {x : E} {n : ℕ∞} +variables [normed_add_comm_group E] [normed_space ℝ E] [normed_add_comm_group X] [normed_space ℝ X] +[has_cont_diff_bump E] {c : E} (f : cont_diff_bump c) {x : E} {n : ℕ∞} -/-- The function defined by `f : cont_diff_bump_of_inner c`. Use automatic coercion to +/-- The function defined by `f : cont_diff_bump c`. Use automatic coercion to function instead. -/ -def to_fun (f : cont_diff_bump_of_inner c) : E → ℝ := -λ x, real.smooth_transition ((f.R - dist x c) / (f.R - f.r)) +def to_fun {c : E} (f : cont_diff_bump c) : E → ℝ := +λ x, (some_cont_diff_bump_base E).to_fun (f.R / f.r) (f.r⁻¹ • (x - c)) -instance : has_coe_to_fun (cont_diff_bump_of_inner c) (λ _, E → ℝ) := ⟨to_fun⟩ +instance : has_coe_to_fun (cont_diff_bump c) (λ _, E → ℝ) := ⟨to_fun⟩ -protected lemma «def» (x : E) : f x = real.smooth_transition ((f.R - dist x c) / (f.R - f.r)) := +protected lemma «def» (x : E) : + f x = (some_cont_diff_bump_base E).to_fun (f.R / f.r) (f.r⁻¹ • (x - c)) := rfl protected lemma sub (x : E) : f (c - x) = f (c + x) := -by simp_rw [f.def, dist_self_sub_left, dist_self_add_left] +by simp [f.def, cont_diff_bump_base.symmetric] -protected lemma neg (f : cont_diff_bump_of_inner (0 : E)) (x : E) : f (- x) = f x := +protected lemma neg (f : cont_diff_bump (0 : E)) (x : E) : f (- x) = f x := by simp_rw [← zero_sub, f.sub, zero_add] -open real (smooth_transition) real.smooth_transition metric +open metric lemma one_of_mem_closed_ball (hx : x ∈ closed_ball c f.r) : f x = 1 := -one_of_one_le $ (one_le_div (sub_pos.2 f.r_lt_R)).2 $ sub_le_sub_left hx _ +begin + apply cont_diff_bump_base.eq_one _ _ f.one_lt_R_div_r, + simpa only [norm_smul, norm_eq_abs, abs_inv, abs_of_nonneg f.r_pos.le, ← div_eq_inv_mul, + div_le_one f.r_pos] using mem_closed_ball_iff_norm.1 hx +end -lemma nonneg : 0 ≤ f x := nonneg _ +lemma nonneg : 0 ≤ f x := +(cont_diff_bump_base.mem_Icc ((some_cont_diff_bump_base E)) _ _).1 -/-- A version of `cont_diff_bump_of_inner.nonneg` with `x` explicit -/ +/-- A version of `cont_diff_bump.nonneg` with `x` explicit -/ lemma nonneg' (x : E) : 0 ≤ f x := f.nonneg -lemma le_one : f x ≤ 1 := le_one _ +lemma le_one : f x ≤ 1 := +(cont_diff_bump_base.mem_Icc ((some_cont_diff_bump_base E)) _ _).2 lemma pos_of_mem_ball (hx : x ∈ ball c f.R) : 0 < f x := -pos_of_pos $ div_pos (sub_pos.2 hx) (sub_pos.2 f.r_lt_R) - -lemma lt_one_of_lt_dist (h : f.r < dist x c) : f x < 1 := -lt_one_of_lt_one $ (div_lt_one (sub_pos.2 f.r_lt_R)).2 $ sub_lt_sub_left h _ +begin + refine lt_iff_le_and_ne.2 ⟨f.nonneg, ne.symm _⟩, + change (f.r)⁻¹ • (x - c) ∈ support ((some_cont_diff_bump_base E).to_fun (f.R / f.r)), + rw cont_diff_bump_base.support _ _ f.one_lt_R_div_r, + simp only [dist_eq_norm, mem_ball] at hx, + simpa only [norm_smul, mem_ball_zero_iff, norm_eq_abs, abs_inv, abs_of_nonneg f.r_pos.le, + ← div_eq_inv_mul] using (div_lt_div_right f.r_pos).2 hx, +end lemma zero_of_le_dist (hx : f.R ≤ dist x c) : f x = 0 := -zero_of_nonpos $ div_nonpos_of_nonpos_of_nonneg (sub_nonpos.2 hx) (sub_nonneg.2 f.r_lt_R.le) +begin + rw dist_eq_norm at hx, + suffices H : (f.r)⁻¹ • (x - c) ∉ support ((some_cont_diff_bump_base E).to_fun (f.R / f.r)), + by simpa only [mem_support, not_not] using H, + rw cont_diff_bump_base.support _ _ f.one_lt_R_div_r, + simp [norm_smul, norm_eq_abs, abs_inv, abs_of_nonneg f.r_pos.le, ← div_eq_inv_mul], + exact div_le_div_of_le f.r_pos.le hx, +end lemma support_eq : support (f : E → ℝ) = metric.ball c f.R := begin ext x, suffices : f x ≠ 0 ↔ dist x c < f.R, by simpa [mem_support], cases lt_or_le (dist x c) f.R with hx hx, - { simp [hx, (f.pos_of_mem_ball hx).ne'] }, - { simp [hx.not_lt, f.zero_of_le_dist hx] } + { simp only [hx, (f.pos_of_mem_ball hx).ne', ne.def, not_false_iff]}, + { simp only [hx.not_lt, f.zero_of_le_dist hx, ne.def, eq_self_iff_true, not_true] } end lemma tsupport_eq : tsupport f = closed_ball c f.R := @@ -363,7 +450,7 @@ f.eventually_eq_one_of_mem_ball (mem_ball_self f.r_pos) /-- `cont_diff_bump` is `𝒞ⁿ` in all its arguments. -/ protected lemma _root_.cont_diff_at.cont_diff_bump {c g : X → E} - {f : ∀ x, cont_diff_bump_of_inner (c x)} {x : X} + {f : ∀ x, cont_diff_bump (c x)} {x : X} (hc : cont_diff_at ℝ n c x) (hr : cont_diff_at ℝ n (λ x, (f x).r) x) (hR : cont_diff_at ℝ n (λ x, (f x).R) x) (hg : cont_diff_at ℝ n g x) : cont_diff_at ℝ n (λ x, f x (g x)) x := @@ -376,11 +463,18 @@ begin exact eventually_of_mem this (λ x hx, (f x).one_of_mem_closed_ball (mem_set_of_eq.mp hx).le) }, exact cont_diff_at_const.congr_of_eventually_eq this }, - { refine real.smooth_transition.cont_diff_at.comp x _, - refine ((hR.sub $ hg.dist hc hx).div (hR.sub hr) (sub_pos.mpr (f x).r_lt_R).ne') } + { change cont_diff_at ℝ n ((uncurry (some_cont_diff_bump_base E).to_fun) ∘ + (λ (x : X), ((f x).R / (f x).r, ((f x).r)⁻¹ • (g x - c x)))) x, + have A : ((f x).R / (f x).r, ((f x).r)⁻¹ • (g x - c x)) ∈ Ioi (1 : ℝ) ×ˢ (univ : set E), + by simpa only [prod_mk_mem_set_prod_eq, mem_univ, and_true] using (f x).one_lt_R_div_r, + have B : Ioi (1 : ℝ) ×ˢ (univ : set E) ∈ 𝓝 ((f x).R / (f x).r, (f x).r⁻¹ • (g x - c x)), + from (is_open_Ioi.prod is_open_univ).mem_nhds A, + apply ((((some_cont_diff_bump_base E).smooth.cont_diff_within_at A).cont_diff_at B) + .of_le le_top).comp x _, + exact (hR.div hr (f x).r_pos.ne').prod ((hr.inv (f x).r_pos.ne').smul (hg.sub hc)) } end -lemma _root_.cont_diff.cont_diff_bump {c g : X → E} {f : ∀ x, cont_diff_bump_of_inner (c x)} +lemma _root_.cont_diff.cont_diff_bump {c g : X → E} {f : ∀ x, cont_diff_bump (c x)} (hc : cont_diff ℝ n c) (hr : cont_diff ℝ n (λ x, (f x).r)) (hR : cont_diff ℝ n (λ x, (f x).R)) (hg : cont_diff ℝ n g) : cont_diff ℝ n (λ x, f x (g x)) := by { rw [cont_diff_iff_cont_diff_at] at *, exact λ x, (hc x).cont_diff_bump (hr x) (hR x) (hg x) } @@ -419,7 +513,7 @@ f.continuous.div_const lemma normed_sub (x : E) : f.normed μ (c - x) = f.normed μ (c + x) := by simp_rw [f.normed_def, f.sub] -lemma normed_neg (f : cont_diff_bump_of_inner (0 : E)) (x : E) : f.normed μ (- x) = f.normed μ x := +lemma normed_neg (f : cont_diff_bump (0 : E)) (x : E) : f.normed μ (- x) = f.normed μ x := by simp_rw [f.normed_def, f.neg] variables [borel_space E] [finite_dimensional ℝ E] [is_locally_finite_measure μ] @@ -441,13 +535,13 @@ end lemma integral_normed : ∫ x, f.normed μ x ∂μ = 1 := begin - simp_rw [cont_diff_bump_of_inner.normed, div_eq_mul_inv, mul_comm (f _), ← smul_eq_mul, + simp_rw [cont_diff_bump.normed, div_eq_mul_inv, mul_comm (f _), ← smul_eq_mul, integral_smul], exact inv_mul_cancel (f.integral_pos.ne') end lemma support_normed_eq : support (f.normed μ) = metric.ball c f.R := -by simp_rw [cont_diff_bump_of_inner.normed, support_div, f.support_eq, +by simp_rw [cont_diff_bump.normed, support_div, f.support_eq, support_const f.integral_pos.ne', inter_univ] lemma tsupport_normed_eq : tsupport (f.normed μ) = metric.closed_ball c f.R := @@ -456,7 +550,7 @@ by simp_rw [tsupport, f.support_normed_eq, closure_ball _ f.R_pos.ne'] lemma has_compact_support_normed : has_compact_support (f.normed μ) := by simp_rw [has_compact_support, f.tsupport_normed_eq, is_compact_closed_ball] -lemma tendsto_support_normed_small_sets {ι} {φ : ι → cont_diff_bump_of_inner c} {l : filter ι} +lemma tendsto_support_normed_small_sets {ι} {φ : ι → cont_diff_bump c} {l : filter ι} (hφ : tendsto (λ i, (φ i).R) l (𝓝 0)) : tendsto (λ i, support (λ x, (φ i).normed μ x)) l (𝓝 c).small_sets := begin @@ -471,118 +565,8 @@ begin end variable (μ) -lemma integral_normed_smul (z : X) [complete_space X] : ∫ x, f.normed μ x • z ∂μ = z := +lemma integral_normed_smul [complete_space X] (z : X) : + ∫ x, f.normed μ x • z ∂μ = z := by simp_rw [integral_smul_const, f.integral_normed, one_smul] -end cont_diff_bump_of_inner - -/-- `f : cont_diff_bump c`, where `c` is a point in a finite dimensional real vector space, is -a bundled smooth function such that - - - `f` is equal to `1` in `euclidean.closed_ball c f.r`; - - `support f = euclidean.ball c f.R`; - - `0 ≤ f x ≤ 1` for all `x`. - -The structure `cont_diff_bump` contains the data required to construct the function: real -numbers `r`, `R`, and proofs of `0 < r < R`. The function itself is available through `coe_fn`.-/ -structure cont_diff_bump [normed_add_comm_group E] [normed_space ℝ E] [finite_dimensional ℝ E] - (c : E) - extends cont_diff_bump_of_inner (to_euclidean c) - -namespace cont_diff_bump - -variables [normed_add_comm_group E] [normed_space ℝ E] [finite_dimensional ℝ E] {c x : E} - (f : cont_diff_bump c) - -/-- The function defined by `f : cont_diff_bump c`. Use automatic coercion to function -instead. -/ -def to_fun (f : cont_diff_bump c) : E → ℝ := f.to_cont_diff_bump_of_inner ∘ to_euclidean - -instance : has_coe_to_fun (cont_diff_bump c) (λ _, E → ℝ) := ⟨to_fun⟩ - -instance (c : E) : inhabited (cont_diff_bump c) := ⟨⟨default⟩⟩ - -lemma R_pos : 0 < f.R := f.to_cont_diff_bump_of_inner.R_pos - -lemma coe_eq_comp : ⇑f = f.to_cont_diff_bump_of_inner ∘ to_euclidean := rfl - -lemma one_of_mem_closed_ball (hx : x ∈ euclidean.closed_ball c f.r) : - f x = 1 := -f.to_cont_diff_bump_of_inner.one_of_mem_closed_ball hx - -lemma nonneg : 0 ≤ f x := f.to_cont_diff_bump_of_inner.nonneg - -lemma le_one : f x ≤ 1 := f.to_cont_diff_bump_of_inner.le_one - -lemma pos_of_mem_ball (hx : x ∈ euclidean.ball c f.R) : 0 < f x := -f.to_cont_diff_bump_of_inner.pos_of_mem_ball hx - -lemma lt_one_of_lt_dist (h : f.r < euclidean.dist x c) : f x < 1 := -f.to_cont_diff_bump_of_inner.lt_one_of_lt_dist h - -lemma zero_of_le_dist (hx : f.R ≤ euclidean.dist x c) : f x = 0 := -f.to_cont_diff_bump_of_inner.zero_of_le_dist hx - -lemma support_eq : support (f : E → ℝ) = euclidean.ball c f.R := -by rw [euclidean.ball_eq_preimage, ← f.to_cont_diff_bump_of_inner.support_eq, - ← support_comp_eq_preimage, coe_eq_comp] - -lemma tsupport_eq : tsupport f = euclidean.closed_ball c f.R := -by rw [tsupport, f.support_eq, euclidean.closure_ball _ f.R_pos.ne'] - -protected lemma has_compact_support : has_compact_support f := -by simp_rw [has_compact_support, f.tsupport_eq, euclidean.is_compact_closed_ball] - -lemma eventually_eq_one_of_mem_ball (h : x ∈ euclidean.ball c f.r) : - f =ᶠ[𝓝 x] 1 := -to_euclidean.continuous_at (f.to_cont_diff_bump_of_inner.eventually_eq_one_of_mem_ball h) - -lemma eventually_eq_one : f =ᶠ[𝓝 c] 1 := -f.eventually_eq_one_of_mem_ball $ euclidean.mem_ball_self f.r_pos - -protected lemma cont_diff {n} : - cont_diff ℝ n f := -f.to_cont_diff_bump_of_inner.cont_diff.comp (to_euclidean : E ≃L[ℝ] _).cont_diff - -protected lemma cont_diff_at {n} : - cont_diff_at ℝ n f x := -f.cont_diff.cont_diff_at - -protected lemma cont_diff_within_at {s n} : - cont_diff_within_at ℝ n f s x := -f.cont_diff_at.cont_diff_within_at - -lemma exists_tsupport_subset {s : set E} (hs : s ∈ 𝓝 c) : - ∃ f : cont_diff_bump c, tsupport f ⊆ s := -let ⟨R, h0, hR⟩ := euclidean.nhds_basis_closed_ball.mem_iff.1 hs -in ⟨⟨⟨R / 2, R, half_pos h0, half_lt_self h0⟩⟩, by rwa tsupport_eq⟩ - -lemma exists_closure_subset {R : ℝ} (hR : 0 < R) - {s : set E} (hs : is_closed s) (hsR : s ⊆ euclidean.ball c R) : - ∃ f : cont_diff_bump c, f.R = R ∧ s ⊆ euclidean.ball c f.r := -begin - rcases euclidean.exists_pos_lt_subset_ball hR hs hsR with ⟨r, hr, hsr⟩, - exact ⟨⟨⟨r, R, hr.1, hr.2⟩⟩, rfl, hsr⟩ -end - end cont_diff_bump - -open finite_dimensional metric - -/-- If `E` is a finite dimensional normed space over `ℝ`, then for any point `x : E` and its -neighborhood `s` there exists an infinitely smooth function with the following properties: - -* `f y = 1` in a neighborhood of `x`; -* `f y = 0` outside of `s`; -* moreover, `tsupport f ⊆ s` and `f` has compact support; -* `f y ∈ [0, 1]` for all `y`. - -This lemma is a simple wrapper around lemmas about bundled smooth bump functions, see -`cont_diff_bump`. -/ -lemma exists_cont_diff_bump_function_of_mem_nhds [normed_add_comm_group E] [normed_space ℝ E] - [finite_dimensional ℝ E] {x : E} {s : set E} (hs : s ∈ 𝓝 x) : - ∃ f : E → ℝ, f =ᶠ[𝓝 x] 1 ∧ (∀ y, f y ∈ Icc (0 : ℝ) 1) ∧ cont_diff ℝ ⊤ f ∧ - has_compact_support f ∧ tsupport f ⊆ s := -let ⟨f, hf⟩ := cont_diff_bump.exists_tsupport_subset hs in -⟨f, f.eventually_eq_one, λ y, ⟨f.nonneg, f.le_one⟩, f.cont_diff, - f.has_compact_support, hf⟩ diff --git a/src/analysis/convolution.lean b/src/analysis/convolution.lean index f21b490b15f2e..be505012c1379 100644 --- a/src/analysis/convolution.lean +++ b/src/analysis/convolution.lean @@ -6,8 +6,8 @@ Authors: Floris van Doorn import measure_theory.group.integration import measure_theory.group.prod import measure_theory.function.locally_integrable +import analysis.calculus.bump_function_inner import measure_theory.integral.interval_integral -import analysis.calculus.specific_functions import analysis.calculus.parametric_integral /-! @@ -66,7 +66,7 @@ Versions of these statements for functions depending on a parameter are also giv * `convolution_tendsto_right`: Given a sequence of nonnegative normalized functions whose support tends to a small neighborhood around `0`, the convolution tends to the right argument. - This is specialized to bump functions in `cont_diff_bump_of_inner.convolution_tendsto_right`. + This is specialized to bump functions in `cont_diff_bump.convolution_tendsto_right`. # Notation The following notations are localized in the locale `convolution`: @@ -902,7 +902,7 @@ end * `g i x` tends to `z₀` as `(i, x)` tends to `l ×ᶠ 𝓝 x₀`; * `k i` tends to `x₀`. -See also `cont_diff_bump_of_inner.convolution_tendsto_right`. +See also `cont_diff_bump.convolution_tendsto_right`. -/ lemma convolution_tendsto_right {ι} {g : ι → G → E'} {l : filter ι} {x₀ : G} {z₀ : E'} @@ -937,13 +937,13 @@ end end normed_add_comm_group -namespace cont_diff_bump_of_inner +namespace cont_diff_bump variables {n : ℕ∞} variables [normed_space ℝ E'] -variables [inner_product_space ℝ G] +variables [normed_add_comm_group G] [normed_space ℝ G] [has_cont_diff_bump G] variables [complete_space E'] -variables {a : G} {φ : cont_diff_bump_of_inner (0 : G)} +variables {a : G} {φ : cont_diff_bump (0 : G)} /-- If `φ` is a bump function, compute `(φ ⋆ g) x₀` if `g` is constant on `metric.ball x₀ φ.R`. -/ lemma convolution_eq_right {x₀ : G} @@ -976,7 +976,7 @@ dist_convolution_le (by simp_rw [← dist_self (g x₀), hg x₀ (mem_ball_self * `g i` is `mu`-a.e. strongly measurable as `i` tends to `l`; * `g i x` tends to `z₀` as `(i, x)` tends to `l ×ᶠ 𝓝 x₀`; * `k i` tends to `x₀`. -/ -lemma convolution_tendsto_right {ι} {φ : ι → cont_diff_bump_of_inner (0 : G)} +lemma convolution_tendsto_right {ι} {φ : ι → cont_diff_bump (0 : G)} {g : ι → G → E'} {k : ι → G} {x₀ : G} {z₀ : E'} {l : filter ι} (hφ : tendsto (λ i, (φ i).R) l (𝓝 0)) (hig : ∀ᶠ i in l, ae_strongly_measurable (g i) μ) @@ -987,16 +987,16 @@ convolution_tendsto_right (eventually_of_forall $ λ i, (φ i).nonneg_normed) (eventually_of_forall $ λ i, (φ i).integral_normed) (tendsto_support_normed_small_sets hφ) hig hcg hk -/-- Special case of `cont_diff_bump_of_inner.convolution_tendsto_right` where `g` is continuous, +/-- Special case of `cont_diff_bump.convolution_tendsto_right` where `g` is continuous, and the limit is taken only in the first function. -/ -lemma convolution_tendsto_right_of_continuous {ι} {φ : ι → cont_diff_bump_of_inner (0 : G)} +lemma convolution_tendsto_right_of_continuous {ι} {φ : ι → cont_diff_bump (0 : G)} {l : filter ι} (hφ : tendsto (λ i, (φ i).R) l (𝓝 0)) (hg : continuous g) (x₀ : G) : tendsto (λ i, ((λ x, (φ i).normed μ x) ⋆[lsmul ℝ ℝ, μ] g : G → E') x₀) l (𝓝 (g x₀)) := convolution_tendsto_right hφ (eventually_of_forall $ λ _, hg.ae_strongly_measurable) ((hg.tendsto x₀).comp tendsto_snd) tendsto_const_nhds -end cont_diff_bump_of_inner +end cont_diff_bump end measurability diff --git a/src/geometry/manifold/bump_function.lean b/src/geometry/manifold/bump_function.lean index 999bcb4927999..80262412fa133 100644 --- a/src/geometry/manifold/bump_function.lean +++ b/src/geometry/manifold/bump_function.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ -import analysis.calculus.specific_functions +import analysis.calculus.bump_function_findim import geometry.manifold.cont_mdiff /-! @@ -14,8 +14,8 @@ In this file we define `smooth_bump_function I c` to be a bundled smooth "bump" define a coercion to function for this type, and for `f : smooth_bump_function I c`, the function `⇑f` written in the extended chart at `c` has the following properties: -* `f x = 1` in the closed euclidean ball of radius `f.r` centered at `c`; -* `f x = 0` outside of the euclidean ball of radius `f.R` centered at `c`; +* `f x = 1` in the closed ball of radius `f.r` centered at `c`; +* `f x = 0` outside of the ball of radius `f.R` centered at `c`; * `0 ≤ f x ≤ 1` for all `x`. The actual statements involve (pre)images under `ext_chart_at I f` and are given as lemmas in the @@ -32,7 +32,7 @@ variables {H : Type uH} [topological_space H] (I : model_with_corners ℝ E H) {M : Type uM} [topological_space M] [charted_space H M] [smooth_manifold_with_corners I M] -open function filter finite_dimensional set +open function filter finite_dimensional set metric open_locale topology manifold classical filter big_operators noncomputable theory @@ -47,8 +47,8 @@ In this section we define a structure for a bundled smooth bump function and pro `f : smooth_bump_function I M` is a smooth function on `M` such that in the extended chart `e` at `f.c`: -* `f x = 1` in the closed euclidean ball of radius `f.r` centered at `f.c`; -* `f x = 0` outside of the euclidean ball of radius `f.R` centered at `f.c`; +* `f x = 1` in the closed ball of radius `f.r` centered at `f.c`; +* `f x = 0` outside of the ball of radius `f.R` centered at `f.c`; * `0 ≤ f x ≤ 1` for all `x`. The structure contains data required to construct a function with these properties. The function is @@ -56,15 +56,12 @@ available as `⇑f` or `f x`. Formal statements of the properties listed above i (pre)images under `ext_chart_at I f.c` and are given as lemmas in the `smooth_bump_function` namespace. -/ structure smooth_bump_function (c : M) extends cont_diff_bump (ext_chart_at I c c) := -(closed_ball_subset : - (euclidean.closed_ball (ext_chart_at I c c) R) ∩ range I ⊆ (ext_chart_at I c).target) +(closed_ball_subset : (closed_ball (ext_chart_at I c c) R) ∩ range I ⊆ (ext_chart_at I c).target) variable {M} namespace smooth_bump_function -open euclidean (renaming dist -> eudist) - variables {c : M} (f : smooth_bump_function I c) {x : M} {I} /-- The function defined by `f : smooth_bump_function c`. Use automatic coercion to function @@ -93,7 +90,7 @@ lemma eventually_eq_of_mem_source (hx : x ∈ (chart_at H c).source) : f.eq_on_source.eventually_eq_of_mem $ is_open.mem_nhds (chart_at H c).open_source hx lemma one_of_dist_le (hs : x ∈ (chart_at H c).source) - (hd : eudist (ext_chart_at I c x) (ext_chart_at I c c) ≤ f.r) : + (hd : dist (ext_chart_at I c x) (ext_chart_at I c c) ≤ f.r) : f x = 1 := by simp only [f.eq_on_source hs, (∘), f.to_cont_diff_bump.one_of_mem_closed_ball hd] @@ -149,7 +146,7 @@ lemma nonneg : 0 ≤ f x := f.mem_Icc.1 lemma le_one : f x ≤ 1 := f.mem_Icc.2 lemma eventually_eq_one_of_dist_lt (hs : x ∈ (chart_at H c).source) - (hd : eudist (ext_chart_at I c x) (ext_chart_at I c c) < f.r) : + (hd : dist (ext_chart_at I c x) (ext_chart_at I c c) < f.r) : f =ᶠ[𝓝 x] 1 := begin filter_upwards [is_open.mem_nhds (is_open_ext_chart_at_preimage I c is_open_ball) ⟨hs, hd⟩], @@ -159,7 +156,7 @@ end lemma eventually_eq_one : f =ᶠ[𝓝 c] 1 := f.eventually_eq_one_of_dist_lt (mem_chart_source _ _) $ -by { rw [euclidean.dist, dist_self], exact f.r_pos } +by { rw [dist_self], exact f.r_pos } @[simp] lemma eq_one : f c = 1 := f.eventually_eq_one.eq_of_nhds @@ -173,9 +170,9 @@ lemma c_mem_support : c ∈ support f := mem_of_mem_nhds f.support_mem_nhds lemma nonempty_support : (support f).nonempty := ⟨c, f.c_mem_support⟩ -lemma compact_symm_image_closed_ball : +lemma is_compact_symm_image_closed_ball : is_compact ((ext_chart_at I c).symm '' (closed_ball (ext_chart_at I c c) f.R ∩ range I)) := -(euclidean.is_compact_closed_ball.inter_right I.closed_range).image_of_continuous_on $ +((is_compact_closed_ball _ _).inter_right I.closed_range).image_of_continuous_on $ (continuous_on_ext_chart_at_symm _ _).mono f.closed_ball_subset /-- Given a smooth bump function `f : smooth_bump_function I c`, the closed ball of radius `f.R` is @@ -185,11 +182,11 @@ lemma nhds_within_range_basis : (𝓝[range I] (ext_chart_at I c c)).has_basis (λ f : smooth_bump_function I c, true) (λ f, closed_ball (ext_chart_at I c c) f.R ∩ range I) := begin - refine ((nhds_within_has_basis euclidean.nhds_basis_closed_ball _).restrict_subset + refine ((nhds_within_has_basis nhds_basis_closed_ball _).restrict_subset (ext_chart_at_target_mem_nhds_within _ _)).to_has_basis' _ _, { rintro R ⟨hR0, hsub⟩, - exact ⟨⟨⟨⟨R / 2, R, half_pos hR0, half_lt_self hR0⟩⟩, hsub⟩, trivial, subset.rfl⟩ }, - { exact λ f _, inter_mem (mem_nhds_within_of_mem_nhds $ closed_ball_mem_nhds f.R_pos) + exact ⟨⟨⟨R / 2, R, half_pos hR0, half_lt_self hR0⟩, hsub⟩, trivial, subset.rfl⟩ }, + { exact λ f _, inter_mem (mem_nhds_within_of_mem_nhds $ closed_ball_mem_nhds _ f.R_pos) self_mem_nhds_within } end @@ -199,7 +196,7 @@ begin rw f.image_eq_inter_preimage_of_subset_support hs, refine continuous_on.preimage_closed_of_closed ((continuous_on_ext_chart_at_symm _ _).mono f.closed_ball_subset) _ hsc, - exact is_closed.inter is_closed_closed_ball I.closed_range + exact is_closed.inter is_closed_ball I.closed_range end /-- If `f` is a smooth bump function and `s` closed subset of the support of `f` (i.e., of the open @@ -212,13 +209,13 @@ begin set e := ext_chart_at I c, have : is_closed (e '' s) := f.is_closed_image_of_is_closed hsc hs, rw [support_eq_inter_preimage, subset_inter_iff, ← image_subset_iff] at hs, - rcases euclidean.exists_pos_lt_subset_ball f.R_pos this hs.2 with ⟨r, hrR, hr⟩, + rcases exists_pos_lt_subset_ball f.R_pos this hs.2 with ⟨r, hrR, hr⟩, exact ⟨r, hrR, subset_inter hs.1 (image_subset_iff.1 hr)⟩ end /-- Replace `r` with another value in the interval `(0, f.R)`. -/ def update_r (r : ℝ) (hr : r ∈ Ioo 0 f.R) : smooth_bump_function I c := -⟨⟨⟨r, f.R, hr.1, hr.2⟩⟩, f.closed_ball_subset⟩ +⟨⟨r, f.R, hr.1, hr.2⟩, f.closed_ball_subset⟩ @[simp] lemma update_r_R {r : ℝ} (hr : r ∈ Ioo 0 f.R) : (f.update_r r hr).R = f.R := rfl @@ -235,7 +232,7 @@ variables [t2_space M] lemma is_closed_symm_image_closed_ball : is_closed ((ext_chart_at I c).symm '' (closed_ball (ext_chart_at I c c) f.R ∩ range I)) := -f.compact_symm_image_closed_ball.is_closed +f.is_compact_symm_image_closed_ball.is_closed lemma tsupport_subset_symm_image_closed_ball : tsupport f ⊆ (ext_chart_at I c).symm '' (closed_ball (ext_chart_at I c c) f.R ∩ range I) := @@ -260,7 +257,7 @@ lemma tsupport_subset_chart_at_source : by simpa only [ext_chart_at_source] using f.tsupport_subset_ext_chart_at_source protected lemma has_compact_support : has_compact_support f := -is_compact_of_is_closed_subset f.compact_symm_image_closed_ball is_closed_closure +is_compact_of_is_closed_subset f.is_compact_symm_image_closed_ball is_closed_closure f.tsupport_subset_symm_image_closed_ball variables (I c) From 69966de79409be171fcfc08e9ea9add087325892 Mon Sep 17 00:00:00 2001 From: Moritz Doll Date: Wed, 8 Feb 2023 10:26:24 +0000 Subject: [PATCH 0438/1291] doc(topology/uniform_space/abstract_completion): typo (#18399) --- src/topology/uniform_space/abstract_completion.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/topology/uniform_space/abstract_completion.lean b/src/topology/uniform_space/abstract_completion.lean index 06110a8d31994..f948d01536302 100644 --- a/src/topology/uniform_space/abstract_completion.lean +++ b/src/topology/uniform_space/abstract_completion.lean @@ -35,7 +35,7 @@ derived from the predicate is more universe polymorphic. ## References We don't know any traditional text discussing this. Real world mathematics simply silently -identify the results of any two constructions that lead to something one could reasonnably +identify the results of any two constructions that lead to something one could reasonably call a completion. ## Tags From 2407f3b12a449410f75ce2877c1f8cb1dc20245d Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 8 Feb 2023 12:29:05 +0000 Subject: [PATCH 0439/1291] feat(analysis/complex/basic): lemmas about tsum (#18376) This adds lemmas about the four "basis" complex operations: `re`, `im`, `of_real` (coercion), and `conj`. These all follow trivially from the API for continuous linear morphisms, but using those results directly gives syntacticly different terms that don't work in rewrites. Similarly, the conj lemmas follow trivially from the results about `star`, but `conj` (aka `star_ring_end`) is not a syntactic match for `star`. This proves all the results for `is_R_or_C` (including a more general version of `has_sum_iff`), then copies across the results to `complex` for convenience. --- src/analysis/complex/basic.lean | 133 +++++++++++++++++++++++++++++--- 1 file changed, 121 insertions(+), 12 deletions(-) diff --git a/src/analysis/complex/basic.lean b/src/analysis/complex/basic.lean index 13819d109a3b3..c83008b39af8b 100644 --- a/src/analysis/complex/basic.lean +++ b/src/analysis/complex/basic.lean @@ -324,22 +324,12 @@ by rw [eq_re_of_real_le hz, is_R_or_C.norm_of_real, real.norm_of_nonneg (complex end complex_order -lemma has_sum_iff {α} (f : α → ℂ) (c : ℂ) : - has_sum f c ↔ has_sum (λ x, (f x).re) c.re ∧ has_sum (λ x, (f x).im) c.im := -begin - -- For some reason, `continuous_linear_map.has_sum` is orders of magnitude faster than - -- `has_sum.mapL` here: - refine ⟨λ h, ⟨re_clm.has_sum h, im_clm.has_sum h⟩, _⟩, - rintro ⟨h₁, h₂⟩, - convert (h₁.prod_mk h₂).mapL equiv_real_prod_clm.symm.to_continuous_linear_map, - { ext x; refl }, - { cases c, refl } -end - end complex namespace is_R_or_C +open_locale complex_conjugate + local notation `reC` := @is_R_or_C.re ℂ _ local notation `imC` := @is_R_or_C.im ℂ _ local notation `IC` := @is_R_or_C.I ℂ _ @@ -354,4 +344,123 @@ by simp [is_R_or_C.norm_sq, complex.norm_sq] @[simp] lemma abs_to_complex {x : ℂ} : absC x = complex.abs x := by simp [is_R_or_C.abs, complex.abs] +section tsum +variables {α : Type*} (𝕜 : Type*) [is_R_or_C 𝕜] + +@[simp] lemma has_sum_conj {f : α → 𝕜} {x : 𝕜} : + has_sum (λ x, conj (f x)) x ↔ has_sum f (conj x) := +conj_cle.has_sum + +lemma has_sum_conj' {f : α → 𝕜} {x : 𝕜} : has_sum (λ x, conj (f x)) (conj x) ↔ has_sum f x := +conj_cle.has_sum' + +@[simp] lemma summable_conj {f : α → 𝕜} : summable (λ x, conj (f x)) ↔ summable f := +summable_star_iff + +variables {𝕜} + +lemma conj_tsum (f : α → 𝕜) : conj (∑' a, f a) = ∑' a, conj (f a) := +tsum_star + +variables (𝕜) + +@[simp, norm_cast] lemma has_sum_of_real {f : α → ℝ} {x : ℝ} : + has_sum (λ x, (f x : 𝕜)) x ↔ has_sum f x := +⟨λ h, by simpa only [is_R_or_C.re_clm_apply, is_R_or_C.of_real_re] using re_clm.has_sum h, + of_real_clm.has_sum⟩ + +@[simp, norm_cast] lemma summable_of_real {f : α → ℝ} : summable (λ x, (f x : 𝕜)) ↔ summable f := +⟨λ h, by simpa only [is_R_or_C.re_clm_apply, is_R_or_C.of_real_re] using re_clm.summable h, + of_real_clm.summable⟩ + +@[norm_cast] lemma of_real_tsum (f : α → ℝ) : (↑(∑' a, f a) : 𝕜) = ∑' a, f a := +begin + by_cases h : summable f, + { exact continuous_linear_map.map_tsum of_real_clm h }, + { rw [tsum_eq_zero_of_not_summable h, + tsum_eq_zero_of_not_summable ((summable_of_real _).not.mpr h), of_real_zero] } +end + +lemma has_sum_re {f : α → 𝕜} {x : 𝕜} (h : has_sum f x) : has_sum (λ x, re (f x)) (re x) := +re_clm.has_sum h + +lemma has_sum_im {f : α → 𝕜} {x : 𝕜} (h : has_sum f x) : has_sum (λ x, im (f x)) (im x) := +im_clm.has_sum h + +lemma re_tsum {f : α → 𝕜} (h : summable f) : re (∑' a, f a) = ∑' a, re (f a) := +re_clm.map_tsum h + +lemma im_tsum {f : α → 𝕜} (h : summable f) : im (∑' a, f a) = ∑' a, im (f a) := +im_clm.map_tsum h + +variables {𝕜} + +lemma has_sum_iff (f : α → 𝕜) (c : 𝕜) : + has_sum f c ↔ has_sum (λ x, re (f x)) (re c) ∧ has_sum (λ x, im (f x)) (im c) := +begin + refine ⟨λ h, ⟨has_sum_re _ h, has_sum_im _ h⟩, _⟩, + rintro ⟨h₁, h₂⟩, + rw ←is_R_or_C.re_add_im c, + convert ((has_sum_of_real 𝕜).mpr h₁).add (((has_sum_of_real 𝕜).mpr h₂).mul_right I), + simp_rw is_R_or_C.re_add_im, +end + +end tsum + end is_R_or_C + +namespace complex +/-! +We have to repeat the lemmas about `is_R_or_C.re` and `is_R_or_C.im` as they are not syntactic +matches for `complex.re` and `complex.im`. + +We do not have this problem with `of_real` and `conj`, although we repeat them anyway for +discoverability and to avoid the need to unify `𝕜`. +-/ +section tsum +variables {α : Type*} + +open_locale complex_conjugate + +@[simp] lemma has_sum_conj {f : α → ℂ} {x : ℂ} : + has_sum (λ x, conj (f x)) x ↔ has_sum f (conj x) := +is_R_or_C.has_sum_conj _ + +lemma has_sum_conj' {f : α → ℂ} {x : ℂ} : has_sum (λ x, conj (f x)) (conj x) ↔ has_sum f x := +is_R_or_C.has_sum_conj' _ + +@[simp] lemma summable_conj {f : α → ℂ} : summable (λ x, conj (f x)) ↔ summable f := +is_R_or_C.summable_conj _ + +lemma conj_tsum (f : α → ℂ) : conj (∑' a, f a) = ∑' a, conj (f a) := +is_R_or_C.conj_tsum _ + +@[simp, norm_cast] lemma has_sum_of_real {f : α → ℝ} {x : ℝ} : + has_sum (λ x, (f x : ℂ)) x ↔ has_sum f x := +is_R_or_C.has_sum_of_real _ + +@[simp, norm_cast] lemma summable_of_real {f : α → ℝ} : summable (λ x, (f x : ℂ)) ↔ summable f := +is_R_or_C.summable_of_real _ + +@[norm_cast] lemma of_real_tsum (f : α → ℝ) : (↑(∑' a, f a) : ℂ) = ∑' a, f a := +is_R_or_C.of_real_tsum _ _ + +lemma has_sum_re {f : α → ℂ} {x : ℂ} (h : has_sum f x) : has_sum (λ x, (f x).re) x.re := +is_R_or_C.has_sum_re _ h + +lemma has_sum_im {f : α → ℂ} {x : ℂ} (h : has_sum f x) : has_sum (λ x, (f x).im) x.im := +is_R_or_C.has_sum_im _ h + +lemma re_tsum {f : α → ℂ} (h : summable f) : (∑' a, f a).re = ∑' a, (f a).re := +is_R_or_C.re_tsum _ h + +lemma im_tsum {f : α → ℂ} (h : summable f) : (∑' a, f a).im = ∑' a, (f a).im := +is_R_or_C.im_tsum _ h + +lemma has_sum_iff (f : α → ℂ) (c : ℂ) : + has_sum f c ↔ has_sum (λ x, (f x).re) c.re ∧ has_sum (λ x, (f x).im) c.im := +is_R_or_C.has_sum_iff _ _ + +end tsum + +end complex From 26b40791e4a5772a4e53d0e28e4df092119dc7da Mon Sep 17 00:00:00 2001 From: prakol16 Date: Wed, 8 Feb 2023 14:04:54 +0000 Subject: [PATCH 0440/1291] feat(combinatorics/catalan): Connection between Catalan numbers and number of trees (#16583) Shows that the number of binary trees with `n` nodes is the `n`th Catalan number. --- src/combinatorics/catalan.lean | 72 +++++++++++++++++++++++++++++++++- 1 file changed, 70 insertions(+), 2 deletions(-) diff --git a/src/combinatorics/catalan.lean b/src/combinatorics/catalan.lean index 177033317785f..b64c70fa5fd44 100644 --- a/src/combinatorics/catalan.lean +++ b/src/combinatorics/catalan.lean @@ -3,9 +3,12 @@ Copyright (c) 2022 Julian Kuelshammer. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Julian Kuelshammer -/ -import data.nat.choose.central import algebra.big_operators.fin +import algebra.big_operators.nat_antidiagonal import algebra.char_zero.lemmas +import data.finset.nat_antidiagonal +import data.nat.choose.central +import data.tree import tactic.field_simp import tactic.linear_combination @@ -26,6 +29,9 @@ triangulations of convex polygons. * `catalan_eq_central_binom_div `: The explicit formula for the Catalan number using the central binomial coefficient, `catalan n = nat.central_binom n / (n + 1)`. +* `trees_of_nodes_eq_card_eq_catalan`: The number of binary trees with `n` internal nodes + is `catalan n` + ## Implementation details The proof of `catalan_eq_central_binom_div` follows @@ -40,7 +46,8 @@ https://math.stackexchange.com/questions/3304415/catalan-numbers-algebraic-proof -/ open_locale big_operators -open finset + +open finset finset.nat.antidiagonal (fst_le snd_le) /-- The recursive definition of the sequence of Catalan numbers: `catalan (n + 1) = ∑ i : fin n.succ, catalan i * catalan (n - i)` -/ @@ -54,6 +61,11 @@ def catalan : ℕ → ℕ lemma catalan_succ (n : ℕ) : catalan (n + 1) = ∑ i : fin n.succ, catalan i * catalan (n - i) := by rw catalan +lemma catalan_succ' (n : ℕ) : + catalan (n + 1) = ∑ ij in nat.antidiagonal n, catalan ij.1 * catalan ij.2 := +by rw [catalan_succ, nat.sum_antidiagonal_eq_sum_range_succ (λ x y, catalan x * catalan y) n, + sum_range] + @[simp] lemma catalan_one : catalan 1 = 1 := by simp [catalan_succ] /-- A helper sequence that can be used to prove the equality of the recursive and the explicit @@ -129,3 +141,59 @@ by norm_num [catalan_eq_central_binom_div, nat.central_binom, nat.choose] lemma catalan_three : catalan 3 = 5 := by norm_num [catalan_eq_central_binom_div, nat.central_binom, nat.choose] + +namespace tree +open_locale tree + +/-- Given two finsets, find all trees that can be formed with + left child in `a` and right child in `b` -/ +@[reducible] def pairwise_node (a b : finset (tree unit)) : finset (tree unit) := +(a ×ˢ b).map ⟨λ x, x.1 △ x.2, λ ⟨x₁, x₂⟩ ⟨y₁, y₂⟩, λ h, by simpa using h⟩ + +/-- A finset of all trees with `n` nodes. See `mem_trees_of_nodes_eq` -/ +def trees_of_num_nodes_eq : ℕ → finset (tree unit) +| 0 := {nil} +| (n+1) := (finset.nat.antidiagonal n).attach.bUnion $ λ ijh, + have _ := nat.lt_succ_of_le (fst_le ijh.2), + have _ := nat.lt_succ_of_le (snd_le ijh.2), + pairwise_node (trees_of_num_nodes_eq ijh.1.1) (trees_of_num_nodes_eq ijh.1.2) + +@[simp] lemma trees_of_nodes_eq_zero : trees_of_num_nodes_eq 0 = {nil} := +by rw [trees_of_num_nodes_eq] + +lemma trees_of_nodes_eq_succ (n : ℕ) : trees_of_num_nodes_eq (n + 1) = + (nat.antidiagonal n).bUnion (λ ij, pairwise_node (trees_of_num_nodes_eq ij.1) + (trees_of_num_nodes_eq ij.2)) := +by { rw trees_of_num_nodes_eq, ext, simp, } + +@[simp] theorem mem_trees_of_nodes_eq {x : tree unit} {n : ℕ} : + x ∈ trees_of_num_nodes_eq n ↔ x.num_nodes = n := +begin + induction x using tree.unit_rec_on generalizing n; + cases n; + simp [trees_of_nodes_eq_succ, nat.succ_eq_add_one, *], + trivial, +end + +lemma mem_trees_of_nodes_eq_num_nodes (x : tree unit) : + x ∈ trees_of_num_nodes_eq x.num_nodes := mem_trees_of_nodes_eq.mpr rfl + +@[simp, norm_cast] lemma coe_trees_of_nodes_eq (n : ℕ) : + ↑(trees_of_num_nodes_eq n) = {x : tree unit | x.num_nodes = n} := set.ext (by simp) + +lemma trees_of_nodes_eq_card_eq_catalan (n : ℕ) : + (trees_of_num_nodes_eq n).card = catalan n := +begin + induction n using nat.case_strong_induction_on with n ih, + { simp, }, + rw [trees_of_nodes_eq_succ, card_bUnion, catalan_succ'], + { apply sum_congr rfl, + rintro ⟨i, j⟩ H, + simp [ih _ (fst_le H), ih _ (snd_le H)], }, + { simp_rw disjoint_left, + rintros ⟨i, j⟩ _ ⟨i', j'⟩ _, + clear_except, + tidy, }, +end + +end tree From bd835ef554f37ef9b804f0903089211f89cb370b Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 8 Feb 2023 15:26:04 +0000 Subject: [PATCH 0441/1291] feat(logic/equiv/fin): `div`/`mod` as an equivalence on `nat`/`int` (#18359) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The motivation here is to be able to use this to change a sum/product/supr/infi/tsum/etc to be indexed by a product, such that every `n`th term can be collected into an inner / outer sum. We already have the API in most places for `equiv`s and `prod`s to do this, all that's missing is the equivalences in this PR. Note that we use `[ne_zero n]` instead of `hn : n ≠ 0` as this is required by `fin.of_nat'` for the coercion. [Zulip](https://leanprover.zulipchat.com/#narrow/stream/217875-Is-there-code-for-X.3F/topic/division.20and.20remainder.20as.20an.20equivalence/near/325359453) Co-authored-by: Junyan Xu --- src/data/int/order/basic.lean | 1 + src/data/nat/basic.lean | 1 + src/logic/equiv/fin.lean | 35 +++++++++++++++++++++++++++++++++++ 3 files changed, 37 insertions(+) diff --git a/src/data/int/order/basic.lean b/src/data/int/order/basic.lean index 94841ad2c1297..05db2a22bca2c 100644 --- a/src/data/int/order/basic.lean +++ b/src/data/int/order/basic.lean @@ -346,6 +346,7 @@ begin rw [sub_add_cancel, ← add_mod_mod, sub_add_cancel, mod_mod] end +/-- See also `int.div_mod_equiv` for a similar statement as an `equiv`. -/ protected theorem div_mod_unique {a b r q : ℤ} (h : 0 < b) : a / b = q ∧ a % b = r ↔ r + b * q = a ∧ 0 ≤ r ∧ r < b := begin diff --git a/src/data/nat/basic.lean b/src/data/nat/basic.lean index d8f47d7509cdd..21a13844f3805 100644 --- a/src/data/nat/basic.lean +++ b/src/data/nat/basic.lean @@ -543,6 +543,7 @@ by { rw mul_comm, exact mod_add_div _ _ } lemma div_add_mod' (m k : ℕ) : (m / k) * k + m % k = m := by { rw mul_comm, exact div_add_mod _ _ } +/-- See also `nat.div_mod_equiv` for a similar statement as an `equiv`. -/ protected theorem div_mod_unique {n k m d : ℕ} (h : 0 < k) : n / k = d ∧ n % k = m ↔ m + k * d = n ∧ m < k := ⟨λ ⟨e₁, e₂⟩, e₁ ▸ e₂ ▸ ⟨mod_add_div _ _, mod_lt _ h⟩, diff --git a/src/logic/equiv/fin.lean b/src/logic/equiv/fin.lean index bc227e87d6552..2033818a7e7ac 100644 --- a/src/logic/equiv/fin.lean +++ b/src/logic/equiv/fin.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau -/ import data.fin.vec_notation +import data.int.order.basic import logic.equiv.defs /-! @@ -424,6 +425,40 @@ def fin_prod_fin_equiv : fin m × fin n ≃ fin (m * n) := ... = y.1 : nat.mod_eq_of_lt y.2), right_inv := λ x, fin.eq_of_veq $ nat.mod_add_div _ _ } +/-- The equivalence induced by `a ↦ (a / n, a % n)` for nonzero `n`. + +This is like `fin_prod_fin_equiv.symm` but with `m` infinite. +See `nat.div_mod_unique` for a similar propositional statement. -/ +@[simps] +def nat.div_mod_equiv (n : ℕ) [ne_zero n] : ℕ ≃ ℕ × fin n := +{ to_fun := λ a, (a / n, ↑a), + inv_fun := λ p, p.1 * n + ↑p.2, -- TODO: is there a canonical order of `*` and `+` here? + left_inv := λ a, nat.div_add_mod' _ _, + right_inv := λ p, begin + refine prod.ext _ (fin.ext $ nat.mul_add_mod_of_lt p.2.is_lt), + dsimp only, + rw [add_comm, nat.add_mul_div_right _ _ (ne_zero.pos n), nat.div_eq_of_lt p.2.is_lt, zero_add], + end } + +/-- The equivalence induced by `a ↦ (a / n, a % n)` for nonzero `n`. + +See `int.div_mod_unique` for a similar propositional statement. -/ +@[simps] +def int.div_mod_equiv (n : ℕ) [ne_zero n] : ℤ ≃ ℤ × fin n := +{ -- TODO: could cast from int directly if we import `data.zmod.defs`, though there are few lemmas + -- about that coercion. + to_fun := λ a, (a / n, ↑(a.nat_mod n)), + inv_fun := λ p, p.1 * n + ↑p.2, + left_inv := λ a, by simp_rw [coe_coe, fin.coe_of_nat_eq_mod, int.coe_nat_mod, int.nat_mod, + int.to_nat_of_nonneg (int.mod_nonneg _ $ ne_zero.ne n), int.mod_mod, int.div_add_mod'], + right_inv := λ ⟨q, r, hrn⟩, begin + simp only [fin.coe_mk, prod.mk.inj_iff, fin.ext_iff, coe_coe], + obtain ⟨h1, h2⟩ := ⟨int.coe_nat_nonneg r, int.coe_nat_lt.2 hrn⟩, + rw [add_comm, int.add_mul_div_right _ _ (ne_zero.ne n), int.div_eq_zero_of_lt h1 h2, + int.nat_mod, int.add_mul_mod_self, int.mod_eq_of_lt h1 h2, int.to_nat_coe_nat], + exact ⟨zero_add q, fin.coe_coe_of_lt hrn⟩, + end } + /-- Promote a `fin n` into a larger `fin m`, as a subtype where the underlying values are retained. This is the `order_iso` version of `fin.cast_le`. -/ @[simps apply symm_apply] From db53863fb135228820ee0b08e8dce9349a3d911b Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9mi=20Bottinelli?= Date: Wed, 8 Feb 2023 18:27:01 +0000 Subject: [PATCH 0442/1291] feat(combinatorics/simple_graph/ends): Definition (only) of the ends of a simple_graph. (#17857) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Anand Rao Co-authored-by: Rémi Bottinelli Co-authored-by: ART0 <18333981+0Art0@users.noreply.github.com> Co-authored-by: Anand Rao <18333981+0art0@users.noreply.github.com> --- src/combinatorics/simple_graph/basic.lean | 23 ++ .../simple_graph/connectivity.lean | 35 +++ src/combinatorics/simple_graph/ends/defs.lean | 238 ++++++++++++++++++ .../simple_graph/ends/properties.lean | 27 ++ 4 files changed, 323 insertions(+) create mode 100644 src/combinatorics/simple_graph/ends/defs.lean create mode 100644 src/combinatorics/simple_graph/ends/properties.lean diff --git a/src/combinatorics/simple_graph/basic.lean b/src/combinatorics/simple_graph/basic.lean index dc4a24b882697..7571e675d0760 100644 --- a/src/combinatorics/simple_graph/basic.lean +++ b/src/combinatorics/simple_graph/basic.lean @@ -1375,6 +1375,29 @@ abbreviation comp (f' : G' ↪g G'') (f : G ↪g G') : G ↪g G'' := f.trans f' end embedding +section induce_hom + +variables {G G'} {G'' : simple_graph X} {s : set V} {t : set W} {r : set X} + (φ : G →g G') (φst : set.maps_to φ s t) (ψ : G' →g G'') (ψtr : set.maps_to ψ t r) + +/-- The restriction of a morphism of graphs to induced subgraphs. -/ +def induce_hom : G.induce s →g G'.induce t := +{ to_fun := set.maps_to.restrict φ s t φst, + map_rel' := λ _ _, φ.map_rel', } + +@[simp, norm_cast] lemma coe_induce_hom : ⇑(induce_hom φ φst) = set.maps_to.restrict φ s t φst := +rfl + +@[simp] lemma induce_hom_id (G : simple_graph V) (s) : + induce_hom (hom.id : G →g G) (set.maps_to_id s) = hom.id := +by { ext x, refl } + +@[simp] lemma induce_hom_comp : + (induce_hom ψ ψtr).comp (induce_hom φ φst) = induce_hom (ψ.comp φ) (ψtr.comp φst) := +by { ext x, refl } + +end induce_hom + namespace iso variables {G G'} (f : G ≃g G') diff --git a/src/combinatorics/simple_graph/connectivity.lean b/src/combinatorics/simple_graph/connectivity.lean index 4a445bd6f9bdd..e0b168cc92085 100644 --- a/src/combinatorics/simple_graph/connectivity.lean +++ b/src/combinatorics/simple_graph/connectivity.lean @@ -877,6 +877,23 @@ end end walk_decomp +/-- +Given a set `S` and a walk `w` from `u` to `v` such that `u ∈ S` but `v ∉ S`, +there exists a dart in the walk whose start is in `S` but whose end is not. +-/ +lemma exists_boundary_dart + {u v : V} (p : G.walk u v) (S : set V) (uS : u ∈ S) (vS : v ∉ S) : + ∃ (d : G.dart), d ∈ p.darts ∧ d.fst ∈ S ∧ d.snd ∉ S := +begin + induction p with _ x y w a p' ih, + { exact absurd uS vS }, + { by_cases h : y ∈ S, + { obtain ⟨d, hd, hcd⟩ := ih h vS, + exact ⟨d, or.inr hd, hcd⟩ }, + { exact ⟨⟨(x, y), a⟩, or.inl rfl, uS, h⟩ } } +end + + end walk /-! ### Type of paths -/ @@ -1484,6 +1501,24 @@ lemma preconnected.subsingleton_connected_component (h : G.preconnected) : subsingleton G.connected_component := ⟨connected_component.ind₂ (λ v w, connected_component.sound (h v w))⟩ +/-- The map on connected components induced by a graph homomorphism. -/ +def connected_component.map {V : Type*} {G : simple_graph V} {V' : Type*} {G' : simple_graph V'} + (φ : G →g G') (C : G.connected_component) : G'.connected_component := +C.lift (λ v, G'.connected_component_mk (φ v)) $ λ v w p _, + connected_component.eq.mpr (p.map φ).reachable + +@[simp] lemma connected_component.map_mk + {V : Type*} {G : simple_graph V} {V' : Type*} {G' : simple_graph V'} (φ : G →g G') (v : V) : + (G.connected_component_mk v).map φ = G'.connected_component_mk (φ v) := rfl + +@[simp] lemma connected_component.map_id (C : connected_component G) : C.map hom.id = C := +by { refine C.ind _, exact (λ _, rfl) } + +@[simp] lemma connected_component.map_comp + {V' : Type*} {G' : simple_graph V'} {V'' : Type*} {G'' : simple_graph V''} + (C : G.connected_component) (φ : G →g G') (ψ : G' →g G'') : (C.map φ).map ψ = C.map (ψ.comp φ) := +by { refine C.ind _, exact (λ _, rfl), } + end connected_component variables {G} diff --git a/src/combinatorics/simple_graph/ends/defs.lean b/src/combinatorics/simple_graph/ends/defs.lean new file mode 100644 index 0000000000000..f86117d54882d --- /dev/null +++ b/src/combinatorics/simple_graph/ends/defs.lean @@ -0,0 +1,238 @@ +/- +Copyright (c) 2022 Anand Rao, Rémi Bottinelli. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Anand Rao, Rémi Bottinelli +-/ +import category_theory.mittag_leffler +import combinatorics.simple_graph.connectivity +import data.set_like.basic + +/-! +# Ends + +This file contains a definition of the ends of a simple graph, as sections of the inverse system +assigning, to each finite set of vertices, the connected components of its complement. +-/ + +universes u +variables {V : Type u} (G : simple_graph V) (K L L' M : set V) + +namespace simple_graph + +/-- The components outside a given set of vertices `K` -/ +@[reducible] def component_compl := (G.induce Kᶜ).connected_component + +variables {G} {K L M} + +/-- The connected component of `v` in `G.induce Kᶜ`. -/ +@[reducible] def component_compl_mk (G : simple_graph V) {v : V} (vK : v ∉ K) : + G.component_compl K := +connected_component_mk (G.induce Kᶜ) ⟨v, vK⟩ + +/-- The set of vertices of `G` making up the connected component `C` -/ +def component_compl.supp (C : G.component_compl K) : set V := +{v : V | ∃ h : v ∉ K, G.component_compl_mk h = C} + +@[ext] lemma component_compl.supp_injective : + function.injective (component_compl.supp : G.component_compl K → set V) := +begin + refine connected_component.ind₂ _, + rintros ⟨v, hv⟩ ⟨w, hw⟩ h, + simp only [set.ext_iff, connected_component.eq, set.mem_set_of_eq, component_compl.supp] at h ⊢, + exact ((h v).mp ⟨hv, reachable.refl _⟩).some_spec, +end + +lemma component_compl.supp_inj {C D : G.component_compl K} : C.supp = D.supp ↔ C = D := +component_compl.supp_injective.eq_iff + +instance : set_like (G.component_compl K) V := +{ coe := component_compl.supp, + coe_injective' := λ C D, (component_compl.supp_inj).mp, } + +@[simp] lemma component_compl.mem_supp_iff {v : V} {C : component_compl G K} : + v ∈ C ↔ ∃ (vK : v ∉ K), G.component_compl_mk vK = C := iff.rfl + +lemma component_compl_mk_mem (G : simple_graph V) {v : V} (vK : v ∉ K) : + v ∈ G.component_compl_mk vK := ⟨vK, rfl⟩ + +lemma component_compl_mk_eq_of_adj (G : simple_graph V) {v w : V} (vK : v ∉ K) (wK : w ∉ K) + (a : G.adj v w) : G.component_compl_mk vK = G.component_compl_mk wK := +by { rw [connected_component.eq], apply adj.reachable, exact a } + +namespace component_compl + +/-- +A `component_compl` specialization of `quot.lift`, where soundness has to be proved only +for adjacent vertices. +-/ +protected def lift {β : Sort*} (f : ∀ ⦃v⦄ (hv : v ∉ K), β) + (h : ∀ ⦃v w⦄ (hv : v ∉ K) (hw : w ∉ K) (a : G.adj v w), f hv = f hw) : G.component_compl K → β := +connected_component.lift (λ vv, f vv.prop) $ (λ v w p, by + { induction p with _ u v w a q ih, + { rintro _, refl, }, + { rintro h', exact (h u.prop v.prop a).trans (ih h'.of_cons), } }) + +protected lemma ind {β : G.component_compl K → Prop} + (f : ∀ ⦃v⦄ (hv : v ∉ K), β (G.component_compl_mk hv)) : ∀ (C : G.component_compl K), β C := by +{ apply connected_component.ind, exact λ ⟨v, vnK⟩, f vnK, } + +/-- The induced graph on the vertices `C`. -/ +@[reducible] +protected def coe_graph (C : component_compl G K) : simple_graph C := G.induce (C : set V) + +lemma coe_inj {C D : G.component_compl K} : (C : set V) = (D : set V) ↔ C = D := set_like.coe_set_eq + +@[simp] protected lemma nonempty (C : G.component_compl K) : (C : set V).nonempty := +C.ind (λ v vnK, ⟨v, vnK, rfl⟩) + +protected lemma exists_eq_mk (C : G.component_compl K) : + ∃ v (h : v ∉ K), G.component_compl_mk h = C := +C.nonempty + +protected lemma disjoint_right (C : G.component_compl K) : disjoint K C := +begin + rw set.disjoint_iff, + exact λ v ⟨vK, vC⟩, vC.some vK, +end + +lemma not_mem_of_mem {C : G.component_compl K} {c : V} (cC : c ∈ C) : c ∉ K := +λ cK, set.disjoint_iff.mp C.disjoint_right ⟨cK, cC⟩ + +protected lemma pairwise_disjoint : + pairwise $ λ C D : G.component_compl K, disjoint (C : set V) (D : set V) := +begin + rintro C D ne, + rw set.disjoint_iff, + exact λ u ⟨uC, uD⟩, ne (uC.some_spec.symm.trans uD.some_spec), +end + +/-- +Any vertex adjacent to a vertex of `C` and not lying in `K` must lie in `C`. +-/ +lemma mem_of_adj : ∀ {C : G.component_compl K} (c d : V), c ∈ C → d ∉ K → G.adj c d → d ∈ C := +λ C c d ⟨cnK, h⟩ dnK cd, + ⟨ dnK, by { rw [←h, connected_component.eq], exact adj.reachable cd.symm, } ⟩ + +/-- +Assuming `G` is preconnected and `K` not empty, given any connected component `C` outside of `K`, +there exists a vertex `k ∈ K` adjacent to a vertex `v ∈ C`. +-/ +lemma exists_adj_boundary_pair (Gc : G.preconnected) (hK : K.nonempty) : + ∀ (C : G.component_compl K), ∃ (ck : V × V), ck.1 ∈ C ∧ ck.2 ∈ K ∧ G.adj ck.1 ck.2 := +begin + refine component_compl.ind (λ v vnK, _), + let C : G.component_compl K := G.component_compl_mk vnK, + let dis := set.disjoint_iff.mp C.disjoint_right, + by_contra' h, + suffices : set.univ = (C : set V), + { exact dis ⟨hK.some_spec, this ▸ (set.mem_univ hK.some)⟩, }, + symmetry, + rw set.eq_univ_iff_forall, + rintro u, + by_contradiction unC, + obtain ⟨p⟩ := Gc v u, + obtain ⟨⟨⟨x, y⟩, xy⟩, d, xC, ynC⟩ := + p.exists_boundary_dart (C : set V) (G.component_compl_mk_mem vnK) unC, + exact ynC (mem_of_adj x y xC (λ (yK : y ∈ K), h ⟨x, y⟩ xC yK xy) xy), +end + +/-- +If `K ⊆ L`, the components outside of `L` are all contained in a single component outside of `K`. +-/ +@[reducible] def hom (h : K ⊆ L) (C : G.component_compl L) : G.component_compl K := +C.map $ induce_hom hom.id $ set.compl_subset_compl.2 h + +lemma subset_hom (C : G.component_compl L) (h : K ⊆ L) : (C : set V) ⊆ (C.hom h : set V) := by +{ rintro c ⟨cL, rfl⟩, exact ⟨λ h', cL (h h'), rfl⟩ } + +lemma _root_.simple_graph.component_compl_mk_mem_hom (G : simple_graph V) {v : V} (vK : v ∉ K) + (h : L ⊆ K) : v ∈ (G.component_compl_mk vK).hom h := +subset_hom (G.component_compl_mk vK) h (G.component_compl_mk_mem vK) + +lemma hom_eq_iff_le (C : G.component_compl L) (h : K ⊆ L) (D : G.component_compl K) : + C.hom h = D ↔ (C : set V) ⊆ (D : set V) := +⟨ λ h', h' ▸ (C.subset_hom h), C.ind (λ v vnL vD, (vD ⟨vnL, rfl⟩).some_spec) ⟩ + +lemma hom_eq_iff_not_disjoint (C : G.component_compl L) (h : K ⊆ L) (D : G.component_compl K) : + C.hom h = D ↔ ¬ disjoint (C : set V) (D : set V) := +begin + rw set.not_disjoint_iff, + split, + { rintro rfl, + apply C.ind (λ x xnL, _), + exact ⟨x, ⟨xnL, rfl⟩, ⟨(λ xK, xnL (h xK)), rfl⟩⟩, }, + { apply C.ind (λ x xnL, _), + rintro ⟨x, ⟨_, e₁⟩, _, rfl⟩, + rw ←e₁, refl, }, +end + +lemma hom_refl (C : G.component_compl L) : C.hom (subset_refl L) = C := +by { change C.map _ = C, erw [induce_hom_id G Lᶜ, connected_component.map_id], } + +lemma hom_trans (C : G.component_compl L) (h : K ⊆ L) (h' : M ⊆ K) : + C.hom (h'.trans h) = (C.hom h).hom h' := +by { change C.map _ = (C.map _).map _, erw [connected_component.map_comp, induce_hom_comp], refl, } + +lemma hom_mk {v : V} (vnL : v ∉ L) (h : K ⊆ L) : + (G.component_compl_mk vnL).hom h = (G.component_compl_mk (set.not_mem_subset h vnL)) := rfl + +lemma hom_infinite (C : G.component_compl L) (h : K ⊆ L) (Cinf : (C : set V).infinite) : + (C.hom h : set V).infinite := set.infinite.mono (C.subset_hom h) Cinf + +lemma infinite_iff_in_all_ranges {K : finset V} (C : G.component_compl K) : + C.supp.infinite ↔ ∀ L (h : K ⊆ L), ∃ D : G.component_compl L, D.hom h = C := +begin + classical, + split, + { rintro Cinf L h, + obtain ⟨v, ⟨vK, rfl⟩, vL⟩ := set.infinite.nonempty (set.infinite.diff Cinf L.finite_to_set), + exact ⟨component_compl_mk _ vL, rfl⟩ }, + { rintro h Cfin, + obtain ⟨D, e⟩ := h (K ∪ Cfin.to_finset) (finset.subset_union_left K Cfin.to_finset), + obtain ⟨v, vD⟩ := D.nonempty, + let Ddis := D.disjoint_right, + simp_rw [finset.coe_union, set.finite.coe_to_finset, set.disjoint_union_left, + set.disjoint_iff] at Ddis, + exact Ddis.right ⟨(component_compl.hom_eq_iff_le _ _ _).mp e vD, vD⟩, }, +end + +end component_compl + +section ends + +variables (G) + +open category_theory + +/-- +The functor assigning, to a finite set in `V`, the set of connected components in its complement. +-/ +@[simps] def component_compl_functor : (finset V)ᵒᵖ ⥤ Type u := +{ obj := λ K, G.component_compl K.unop, + map := λ _ _ f, component_compl.hom (le_of_op_hom f), + map_id' := λ K, funext $ λ C, C.hom_refl, + map_comp' := λ K L M h h', funext $ λ C, C.hom_trans (le_of_op_hom h) (le_of_op_hom h') } + +/-- The end of a graph, defined as the sections of the functor `component_compl_functor` . -/ +@[protected] +def «end» := (component_compl_functor G).sections + +lemma end_hom_mk_of_mk {s} (sec : s ∈ G.end) {K L : (finset V)ᵒᵖ} (h : L ⟶ K) + {v : V} (vnL : v ∉ L.unop) (hs : s L = G.component_compl_mk vnL) : + s K = G.component_compl_mk (set.not_mem_subset (le_of_op_hom h) vnL) := +begin + rw [←(sec h), hs], + apply component_compl.hom_mk, +end + +lemma infinite_iff_in_eventual_range {K : (finset V)ᵒᵖ} (C : G.component_compl_functor.obj K) : + C.supp.infinite ↔ C ∈ G.component_compl_functor.eventual_range K := +begin + simp only [C.infinite_iff_in_all_ranges, category_theory.functor.eventual_range, + set.mem_Inter, set.mem_range, component_compl_functor_map], + exact ⟨λ h Lop KL, h Lop.unop (le_of_op_hom KL), λ h L KL, h (opposite.op L) (op_hom_of_le KL)⟩, +end + +end ends + +end simple_graph diff --git a/src/combinatorics/simple_graph/ends/properties.lean b/src/combinatorics/simple_graph/ends/properties.lean new file mode 100644 index 0000000000000..1daaae2a569e2 --- /dev/null +++ b/src/combinatorics/simple_graph/ends/properties.lean @@ -0,0 +1,27 @@ +/- +Copyright (c) 2022 Anand Rao, Rémi Bottinelli. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Anand Rao, Rémi Bottinelli +-/ +import combinatorics.simple_graph.ends.defs +/-! +# Properties of the ends of graphs + +This file is meant to contain results about the ends of (locally finite connected) graphs. + +-/ + +variables {V : Type} (G : simple_graph V) + +namespace simple_graph + +instance [finite V] : is_empty G.end := +⟨ begin + rintro ⟨s, _⟩, + casesI nonempty_fintype V, + obtain ⟨v, h⟩ := (s $ opposite.op finset.univ).nonempty, + exact set.disjoint_iff.mp (s _).disjoint_right + ⟨by simp only [opposite.unop_op, finset.coe_univ], h⟩, + end ⟩ + +end simple_graph From d101e93197bb5f6ea89bd7ba386b7f7dff1f3903 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Wed, 8 Feb 2023 18:27:03 +0000 Subject: [PATCH 0443/1291] refactor(topology/discrete_quotient): review API (#18401) Backport various changes I made to the API while porting to Lean 4 in leanprover-community/mathlib4#2157. * extend `setoid X`; * require only that the fibers are open in the definition, prove that they are clopen; * golf proofs, reuse lattice structure on `setoid`s; * use bundled continuous maps for `comap`; * swap LHS and RHS in some `simp` lemmas; * generalize the `order_bot` instance from a discrete space to a locally connected space; * prove that a discrete topological space is locally connected. --- src/topology/category/Profinite/as_limit.lean | 7 +- src/topology/connected.lean | 18 +- src/topology/discrete_quotient.lean | 351 ++++++++---------- 3 files changed, 179 insertions(+), 197 deletions(-) diff --git a/src/topology/category/Profinite/as_limit.lean b/src/topology/category/Profinite/as_limit.lean index da4c971533e1a..71d6757804488 100644 --- a/src/topology/category/Profinite/as_limit.lean +++ b/src/topology/category/Profinite/as_limit.lean @@ -40,7 +40,7 @@ variables (X : Profinite.{u}) /-- The functor `discrete_quotient X ⥤ Fintype` whose limit is isomorphic to `X`. -/ def fintype_diagram : discrete_quotient X ⥤ Fintype := -{ obj := λ S, Fintype.of S, +{ obj := λ S, by haveI := fintype.of_finite S; exact Fintype.of S, map := λ S T f, discrete_quotient.of_le f.le } /-- An abbreviation for `X.fintype_diagram ⋙ Fintype_to_Profinite`. -/ @@ -56,9 +56,8 @@ instance is_iso_as_limit_cone_lift : is_iso ((limit_cone_is_limit X.diagram).lift X.as_limit_cone) := is_iso_of_bijective _ begin - refine ⟨λ a b, _, λ a, _⟩, - { intro h, - refine discrete_quotient.eq_of_proj_eq (λ S, _), + refine ⟨λ a b h, _, λ a, _⟩, + { refine discrete_quotient.eq_of_forall_proj_eq (λ S, _), apply_fun (λ f : (limit_cone X.diagram).X, f.val S) at h, exact h }, { obtain ⟨b, hb⟩ := discrete_quotient.exists_of_compat diff --git a/src/topology/connected.lean b/src/topology/connected.lean index b261cc203a9f2..7aab256b6b436 100644 --- a/src/topology/connected.lean +++ b/src/topology/connected.lean @@ -678,6 +678,10 @@ eq_of_subset_of_subset (set.mem_of_mem_of_subset mem_connected_component (is_connected_connected_component.subset_connected_component h))) +theorem connected_component_eq_iff_mem {x y : α} : + connected_component x = connected_component y ↔ x ∈ connected_component y := +⟨λ h, h ▸ mem_connected_component, λ h, (connected_component_eq h).symm⟩ + lemma connected_component_in_eq {x y : α} {F : set α} (h : y ∈ connected_component_in F x) : connected_component_in F x = connected_component_in F y := begin @@ -1168,6 +1172,14 @@ begin λ ⟨V, ⟨hV, hxV, _⟩, hVU⟩, mem_nhds_iff.mpr ⟨V, hVU, hV, hxV⟩⟩⟩ } end +/-- A space with discrete topology is a locally connected space. -/ +@[priority 100] +instance discrete_topology.to_locally_connected_space (α) [topological_space α] + [discrete_topology α] : locally_connected_space α := +locally_connected_space_iff_open_connected_subsets.2 $ λ x _U hU, + ⟨{x}, singleton_subset_iff.2 $ mem_of_mem_nhds hU, is_open_discrete _, mem_singleton _, + is_connected_singleton⟩ + lemma connected_component_in_mem_nhds [locally_connected_space α] {F : set α} {x : α} (h : F ∈ 𝓝 x) : connected_component_in F x ∈ 𝓝 x := @@ -1353,6 +1365,10 @@ begin exact mem_connected_component end +@[simp] theorem connected_component_eq_singleton [totally_disconnected_space α] (x : α) : + connected_component x = {x} := +totally_disconnected_space_iff_connected_component_singleton.1 ‹_› x + /-- The image of a connected component in a totally disconnected space is a singleton. -/ @[simp] lemma continuous.image_connected_component_eq_singleton {β : Type*} [topological_space β] [totally_disconnected_space β] {f : α → β} (h : continuous f) (a : α) : @@ -1463,7 +1479,7 @@ not_congr coe_eq_coe lemma coe_eq_coe' {x y : α} : (x : connected_components α) = y ↔ x ∈ connected_component y := -coe_eq_coe.trans ⟨λ h, h ▸ mem_connected_component, λ h, (connected_component_eq h).symm⟩ +coe_eq_coe.trans connected_component_eq_iff_mem instance [inhabited α] : inhabited (connected_components α) := ⟨↑(default : α)⟩ diff --git a/src/topology/discrete_quotient.lean b/src/topology/discrete_quotient.lean index cdd6ce05a6e50..801bf98c68845 100644 --- a/src/topology/discrete_quotient.lean +++ b/src/topology/discrete_quotient.lean @@ -26,22 +26,31 @@ quotients as setoids whose equivalence classes are clopen. endowed with a `fintype` instance. ## Order structure + The type `discrete_quotient X` is endowed with an instance of a `semilattice_inf` with `order_top`. The partial ordering `A ≤ B` mathematically means that `B.proj` factors through `A.proj`. The top element `⊤` is the trivial quotient, meaning that every element of `X` is collapsed to a point. Given `h : A ≤ B`, the map `A → B` is `discrete_quotient.of_le h`. -Whenever `X` is discrete, the type `discrete_quotient X` is also endowed with an instance of a -`semilattice_inf` with `order_bot`, where the bot element `⊥` is `X` itself. -Given `f : X → Y` and `h : continuous f`, we define a predicate `le_comap h A B` for -`A : discrete_quotient X` and `B : discrete_quotient Y`, asserting that `f` descends to `A → B`. -If `cond : le_comap h A B`, the function `A → B` is obtained by `discrete_quotient.map cond`. +Whenever `X` is a locally connected space, the type `discrete_quotient X` is also endowed with an +instance of a `order_bot`, where the bot element `⊥` is given by the `connectedComponentSetoid`, +i.e., `x ~ y` means that `x` and `y` belong to the same connected component. In particular, if `X` +is a discrete topological space, then `x ~ y` is equivalent (propositionally, not definitionally) to +`x = y`. + +Given `f : C(X, Y)`, we define a predicate `discrete_quotient.le_comap f A B` for `A : +discrete_quotient X` and `B : discrete_quotient Y`, asserting that `f` descends to `A → B`. If +`cond : discrete_quotient.le_comap h A B`, the function `A → B` is obtained by +`discrete_quotient.map f cond`. ## Theorems + The two main results proved in this file are: -1. `discrete_quotient.eq_of_proj_eq` which states that when `X` is compact, t2 and totally - disconnected, any two elements of `X` agree if their projections in `Q` agree for all + +1. `discrete_quotient.eq_of_forall_proj_eq` which states that when `X` is compact, T₂, and totally + disconnected, any two elements of `X` are equal if their projections in `Q` agree for all `Q : discrete_quotient X`. + 2. `discrete_quotient.exists_of_compat` which states that when `X` is compact, then any system of elements of `Q` as `Q : discrete_quotient X` varies, which is compatible with respect to `discrete_quotient.of_le`, must arise from some element of `X`. @@ -51,49 +60,34 @@ The constructions in this file will be used to show that any profinite space is of finite discrete spaces. -/ -variables (X : Type*) [topological_space X] +open set function +variables {α X Y Z : Type*} [topological_space X] [topological_space Y] + [topological_space Z] /-- The type of discrete quotients of a topological space. -/ @[ext] -structure discrete_quotient := -(rel : X → X → Prop) -(equiv : equivalence rel) -(clopen : ∀ x, is_clopen (set_of (rel x))) +structure discrete_quotient (X : Type*) [topological_space X] extends setoid X := +(is_open_set_of_rel : ∀ x, is_open (set_of (to_setoid.rel x))) namespace discrete_quotient -variables {X} (S : discrete_quotient X) +variables (S : discrete_quotient X) /-- Construct a discrete quotient from a clopen set. -/ def of_clopen {A : set X} (h : is_clopen A) : discrete_quotient X := -{ rel := λ x y, x ∈ A ∧ y ∈ A ∨ x ∉ A ∧ y ∉ A, - equiv := ⟨by tauto!, by tauto!, by tauto!⟩, - clopen := begin - intros x, - by_cases hx : x ∈ A, - { apply is_clopen.union, - { convert h, - ext, - exact ⟨λ i, i.2, λ i, ⟨hx,i⟩⟩ }, - { convert is_clopen_empty, - tidy } }, - { apply is_clopen.union, - { convert is_clopen_empty, - tidy }, - { convert is_clopen.compl h, - ext, - exact ⟨λ i, i.2, λ i, ⟨hx, i⟩⟩ } }, - end } - -lemma refl : ∀ x : X, S.rel x x := S.equiv.1 -lemma symm : ∀ x y : X, S.rel x y → S.rel y x := S.equiv.2.1 -lemma trans : ∀ x y z : X, S.rel x y → S.rel y z → S.rel x z := S.equiv.2.2 +{ to_setoid := ⟨λ x y, x ∈ A ↔ y ∈ A, λ _, iff.rfl, λ _ _, iff.symm, λ _ _ _, iff.trans⟩, + is_open_set_of_rel := λ x, + by by_cases hx : x ∈ A; simp [setoid.rel, hx, h.1, h.2, ← compl_set_of] } + +lemma refl : ∀ x, S.rel x x := S.refl' +lemma symm {x y : X} : S.rel x y → S.rel y x := S.symm' +lemma trans {x y z} : S.rel x y → S.rel y z → S.rel x z := S.trans' /-- The setoid whose quotient yields the discrete quotient. -/ -def setoid : setoid X := ⟨S.rel, S.equiv⟩ +add_decl_doc to_setoid instance : has_coe_to_sort (discrete_quotient X) Type* := -⟨λ S, quotient S.setoid⟩ +⟨λ S, quotient S.to_setoid⟩ instance : topological_space S := quotient.topological_space @@ -101,12 +95,7 @@ instance : topological_space S := quotient.topological_space def proj : X → S := quotient.mk' lemma fiber_eq (x : X) : S.proj ⁻¹' {S.proj x} = set_of (S.rel x) := -begin - ext1 y, - simp only [set.mem_preimage, set.mem_singleton_iff, quotient.eq', - discrete_quotient.proj.equations._eqn_1, set.mem_set_of_eq], - exact ⟨λ h, S.symm _ _ h, λ h, S.symm _ _ h⟩, -end +set.ext $ λ y, eq_comm.trans quotient.eq' lemma proj_surjective : function.surjective S.proj := quotient.surjective_quotient_mk' lemma proj_quotient_map : quotient_map S.proj := quotient_map_quot_mk @@ -114,247 +103,225 @@ lemma proj_continuous : continuous S.proj := S.proj_quotient_map.continuous instance : discrete_topology S := singletons_open_iff_discrete.1 $ S.proj_surjective.forall.2 $ λ x, - by { rw [← S.proj_quotient_map.is_open_preimage, fiber_eq], exact (S.clopen _).1 } + by { rw [← S.proj_quotient_map.is_open_preimage, fiber_eq], exact S.is_open_set_of_rel _ } lemma proj_is_locally_constant : is_locally_constant S.proj := (is_locally_constant.iff_continuous S.proj).2 S.proj_continuous -lemma fiber_clopen (A : set S) : is_clopen (S.proj ⁻¹' A) := +lemma is_clopen_preimage (A : set S) : is_clopen (S.proj ⁻¹' A) := (is_clopen_discrete A).preimage S.proj_continuous -lemma fiber_open (A : set S) : is_open (S.proj ⁻¹' A) := (S.fiber_clopen A).1 -lemma fiber_closed (A : set S) : is_closed (S.proj ⁻¹' A) := (S.fiber_clopen A).2 +lemma is_open_preimage (A : set S) : is_open (S.proj ⁻¹' A) := (S.is_clopen_preimage A).1 +lemma is_closed_preimage (A : set S) : is_closed (S.proj ⁻¹' A) := (S.is_clopen_preimage A).2 -instance : partial_order (discrete_quotient X) := -{ le := λ A B, ∀ x y : X, A.rel x y → B.rel x y, - le_refl := λ a, by tauto, - le_trans := λ a b c h1 h2, by tauto, - le_antisymm := λ a b h1 h2, by { ext, tauto } } +theorem is_clopen_set_of_rel (x : X) : is_clopen (set_of (S.rel x)) := +by { rw [← fiber_eq], apply is_clopen_preimage } -instance : order_top (discrete_quotient X) := -{ top := ⟨λ a b, true, ⟨by tauto, by tauto, by tauto⟩, λ _, is_clopen_univ⟩, - le_top := λ a, by tauto } +instance : has_inf (discrete_quotient X) := +⟨λ S₁ S₂, ⟨S₁.1 ⊓ S₂.1, λ x, (S₁.2 x).inter (S₂.2 x)⟩⟩ instance : semilattice_inf (discrete_quotient X) := -{ inf := λ A B, - { rel := λ x y, A.rel x y ∧ B.rel x y, - equiv := ⟨λ a, ⟨A.refl _,B.refl _⟩, λ a b h, ⟨A.symm _ _ h.1, B.symm _ _ h.2⟩, - λ a b c h1 h2, ⟨A.trans _ _ _ h1.1 h2.1, B.trans _ _ _ h1.2 h2.2⟩⟩, - clopen := λ x, is_clopen.inter (A.clopen _) (B.clopen _) }, - inf_le_left := λ a b, by tauto, - inf_le_right := λ a b, by tauto, - le_inf := λ a b c h1 h2, by tauto, - ..discrete_quotient.partial_order } +injective.semilattice_inf to_setoid ext (λ _ _, rfl) + +instance : order_top (discrete_quotient X) := +{ top := ⟨⊤, λ _, is_open_univ⟩, + le_top := λ a, by tauto } instance : inhabited (discrete_quotient X) := ⟨⊤⟩ +instance inhabited_quotient [inhabited X] : inhabited S := ⟨S.proj default⟩ +instance [nonempty X] : nonempty S := nonempty.map S.proj ‹_› + section comap -variables {Y : Type*} [topological_space Y] {f : Y → X} (cont : continuous f) +variables (g : C(Y, Z)) (f : C(X, Y)) /-- Comap a discrete quotient along a continuous map. -/ -def comap : discrete_quotient Y := -{ rel := λ a b, S.rel (f a) (f b), - equiv := ⟨λ a, S.refl _, λ a b h, S.symm _ _ h, λ a b c h1 h2, S.trans _ _ _ h1 h2⟩, - clopen := λ y, ⟨is_open.preimage cont (S.clopen _).1, is_closed.preimage cont (S.clopen _).2⟩ } +def comap (S : discrete_quotient Y) : discrete_quotient X := +{ to_setoid := setoid.comap f S.1, + is_open_set_of_rel := λ y, (S.2 _).preimage f.continuous } @[simp] -lemma comap_id : S.comap (continuous_id : continuous (id : X → X)) = S := by { ext, refl } +lemma comap_id : S.comap (continuous_map.id X) = S := by { ext, refl } @[simp] -lemma comap_comp {Z : Type*} [topological_space Z] {g : Z → Y} (cont' : continuous g) : - S.comap (continuous.comp cont cont') = (S.comap cont).comap cont' := by { ext, refl } +lemma comap_comp (S : discrete_quotient Z) : S.comap (g.comp f) = (S.comap g).comap f := rfl -lemma comap_mono {A B : discrete_quotient X} (h : A ≤ B) : A.comap cont ≤ B.comap cont := +@[mono] +lemma comap_mono {A B : discrete_quotient Y} (h : A ≤ B) : A.comap f ≤ B.comap f := by tauto end comap section of_le +variables {A B C : discrete_quotient X} + /-- The map induced by a refinement of a discrete quotient. -/ -def of_le {A B : discrete_quotient X} (h : A ≤ B) : A → B := -λ a, quotient.lift_on' a (λ x, B.proj x) (λ a b i, quotient.sound' (h _ _ i)) +def of_le (h : A ≤ B) : A → B := quotient.map' (λ x, x) h -@[simp] -lemma of_le_refl {A : discrete_quotient X} : of_le (le_refl A) = id := by { ext ⟨⟩, refl } +@[simp] lemma of_le_refl : of_le (le_refl A) = id := by { ext ⟨⟩, refl } -lemma of_le_refl_apply {A : discrete_quotient X} (a : A) : of_le (le_refl A) a = a := by simp +lemma of_le_refl_apply (a : A) : of_le (le_refl A) a = a := by simp -@[simp] -lemma of_le_comp {A B C : discrete_quotient X} (h1 : A ≤ B) (h2 : B ≤ C) : - of_le (le_trans h1 h2) = of_le h2 ∘ of_le h1 := by { ext ⟨⟩, refl } +@[simp] lemma of_le_of_le (h₁ : A ≤ B) (h₂ : B ≤ C) (x : A) : + of_le h₂ (of_le h₁ x) = of_le (h₁.trans h₂) x := by { rcases x with ⟨⟩, refl } -lemma of_le_comp_apply {A B C : discrete_quotient X} (h1 : A ≤ B) (h2 : B ≤ C) (a : A) : - of_le (le_trans h1 h2) a = of_le h2 (of_le h1 a) := by simp +@[simp] lemma of_le_comp_of_le (h₁ : A ≤ B) (h₂ : B ≤ C) : + of_le h₂ ∘ of_le h₁ = of_le (le_trans h₁ h₂) := +funext $ of_le_of_le _ _ -lemma of_le_continuous {A B : discrete_quotient X} (h : A ≤ B) : - continuous (of_le h) := continuous_of_discrete_topology +lemma of_le_continuous (h : A ≤ B) : continuous (of_le h) := +continuous_of_discrete_topology -@[simp] -lemma of_le_proj {A B : discrete_quotient X} (h : A ≤ B) : - of_le h ∘ A.proj = B.proj := by { ext, exact quotient.sound' (B.refl _) } +@[simp] lemma of_le_proj (h : A ≤ B) (x : X) : of_le h (A.proj x) = B.proj x := +quotient.sound' (B.refl _) -@[simp] -lemma of_le_proj_apply {A B : discrete_quotient X} (h : A ≤ B) (x : X) : - of_le h (A.proj x) = B.proj x := by { change (of_le h ∘ A.proj) x = _, simp } +@[simp] lemma of_le_comp_proj (h : A ≤ B) : of_le h ∘ A.proj = B.proj := +funext $ of_le_proj _ end of_le -/-- -When X is discrete, there is a `order_bot` instance on `discrete_quotient X` +/-- When `X` is a locally connected space, there is an `order_bot` instance on +`discrete_quotient X`. The bottom element is given by `connected_component_setoid X` -/ -instance [discrete_topology X] : order_bot (discrete_quotient X) := +instance [locally_connected_space X] : order_bot (discrete_quotient X) := { bot := - { rel := (=), - equiv := eq_equivalence, - clopen := λ x, is_clopen_discrete _ }, - bot_le := by { rintro S a b (h : a = b), rw h, exact S.refl _ } } - -lemma proj_bot_injective [discrete_topology X] : - function.injective (⊥ : discrete_quotient X).proj := λ a b h, quotient.exact' h - -lemma proj_bot_bijective [discrete_topology X] : - function.bijective (⊥ : discrete_quotient X).proj := ⟨proj_bot_injective, proj_surjective _⟩ + { to_setoid := connected_component_setoid X, + is_open_set_of_rel := λ x, + begin + have : connected_component x = {y | (connected_component_setoid X).rel x y}, + { ext y, + simpa only [connected_component_setoid, ← connected_component_eq_iff_mem] using eq_comm }, + rw [← this], + exact is_open_connected_component + end }, + bot_le := λ S x y (h : connected_component x = connected_component y), + (S.is_clopen_set_of_rel x).connected_component_subset (S.refl _) $ + h.symm ▸ mem_connected_component } + +@[simp] theorem proj_bot_eq [locally_connected_space X] {x y : X} : + proj ⊥ x = proj ⊥ y ↔ connected_component x = connected_component y := +quotient.eq' + +theorem proj_bot_inj [discrete_topology X] {x y : X} : + proj ⊥ x = proj ⊥ y ↔ x = y := by simp + +theorem proj_bot_injective [discrete_topology X] : + injective (⊥ : discrete_quotient X).proj := λ _ _, proj_bot_inj.1 + +theorem proj_bot_bijective [discrete_topology X] : + bijective (⊥ : discrete_quotient X).proj := +⟨proj_bot_injective, proj_surjective _⟩ section map -variables {Y : Type*} [topological_space Y] {f : Y → X} - (cont : continuous f) (A : discrete_quotient Y) (B : discrete_quotient X) +variables (f : C(X, Y)) (A A' : discrete_quotient X) (B B' : discrete_quotient Y) -/-- -Given `cont : continuous f`, `le_comap cont A B` is defined as `A ≤ B.comap f`. -Mathematically this means that `f` descends to a morphism `A → B`. --/ -def le_comap : Prop := A ≤ B.comap cont +/-- Given `f : C(X, Y)`, `le_comap cont A B` is defined as `A ≤ B.comap f`. Mathematically this +means that `f` descends to a morphism `A → B`. -/ +def le_comap : Prop := A ≤ B.comap f + +theorem le_comap_id : le_comap (continuous_map.id X) A A := λ _ _, id -variables {cont A B} +variables {A A' B B'} {f} {g : C(Y, Z)} {C : discrete_quotient Z} -lemma le_comap_id (A : discrete_quotient X) : le_comap continuous_id A A := by tauto +@[simp] theorem le_comap_id_iff : le_comap (continuous_map.id X) A A' ↔ A ≤ A' := iff.rfl -lemma le_comap_comp {Z : Type*} [topological_space Z] {g : Z → Y} {cont' : continuous g} - {C : discrete_quotient Z} : le_comap cont' C A → le_comap cont A B → - le_comap (continuous.comp cont cont') C B := by tauto +theorem le_comap.comp : + le_comap g B C → le_comap f A B → le_comap (g.comp f) A C := by tauto -lemma le_comap_trans {C : discrete_quotient X} : - le_comap cont A B → B ≤ C → le_comap cont A C := λ h1 h2, le_trans h1 $ comap_mono _ h2 +theorem le_comap.mono (h : le_comap f A B) (hA : A' ≤ A) (hB : B ≤ B') : le_comap f A' B' := + hA.trans $ le_trans h $ comap_mono _ hB /-- Map a discrete quotient along a continuous map. -/ -def map (cond : le_comap cont A B) : A → B := quotient.map' f cond +def map (f : C(X, Y)) (cond : le_comap f A B) : A → B := quotient.map' f cond -lemma map_continuous (cond : le_comap cont A B) : continuous (map cond) := -continuous_of_discrete_topology +theorem map_continuous (cond : le_comap f A B) : continuous (map f cond) := + continuous_of_discrete_topology -@[simp] -lemma map_proj (cond : le_comap cont A B) : map cond ∘ A.proj = B.proj ∘ f := rfl +@[simp] theorem map_comp_proj (cond : le_comap f A B) : map f cond ∘ A.proj = B.proj ∘ f := rfl @[simp] -lemma map_proj_apply (cond : le_comap cont A B) (y : Y) : map cond (A.proj y) = B.proj (f y) := rfl +theorem map_proj (cond : le_comap f A B) (x : X) : map f cond (A.proj x) = B.proj (f x) := rfl -@[simp] -lemma map_id : map (le_comap_id A) = id := by { ext ⟨⟩, refl } +@[simp] theorem map_id : map _ (le_comap_id A) = id := by ext ⟨⟩; refl -@[simp] -lemma map_comp {Z : Type*} [topological_space Z] {g : Z → Y} {cont' : continuous g} - {C : discrete_quotient Z} (h1 : le_comap cont' C A) (h2 : le_comap cont A B) : - map (le_comap_comp h1 h2) = map h2 ∘ map h1 := by { ext ⟨⟩, refl } +@[simp] theorem map_comp (h1 : le_comap g B C) (h2 : le_comap f A B) : + map (g.comp f) (h1.comp h2) = map g h1 ∘ map f h2 := +by { ext ⟨⟩, refl } -@[simp] -lemma of_le_map {C : discrete_quotient X} (cond : le_comap cont A B) (h : B ≤ C) : - map (le_comap_trans cond h) = of_le h ∘ map cond := by { ext ⟨⟩, refl } +@[simp] theorem of_le_map (cond : le_comap f A B) (h : B ≤ B') (a : A) : + of_le h (map f cond a) = map f (cond.mono le_rfl h) a := +by { rcases a with ⟨⟩, refl } -@[simp] -lemma of_le_map_apply {C : discrete_quotient X} (cond : le_comap cont A B) (h : B ≤ C) (a : A) : - map (le_comap_trans cond h) a = of_le h (map cond a) := by { rcases a, refl } +@[simp] theorem of_le_comp_map (cond : le_comap f A B) (h : B ≤ B') : + of_le h ∘ map f cond = map f (cond.mono le_rfl h) := +funext $ of_le_map cond h -@[simp] -lemma map_of_le {C : discrete_quotient Y} (cond : le_comap cont A B) (h : C ≤ A) : - map (le_trans h cond) = map cond ∘ of_le h := by { ext ⟨⟩, refl } +@[simp] theorem map_of_le (cond : le_comap f A B) (h : A' ≤ A) (c : A') : + map f cond (of_le h c) = map f (cond.mono h le_rfl) c := +by { rcases c with ⟨⟩, refl } -@[simp] -lemma map_of_le_apply {C : discrete_quotient Y} (cond : le_comap cont A B) (h : C ≤ A) (c : C) : - map (le_trans h cond) c = map cond (of_le h c) := by { rcases c, refl } +@[simp] theorem map_comp_of_le (cond : le_comap f A B) (h : A' ≤ A) : + map f cond ∘ of_le h = map f (cond.mono h le_rfl) := +funext $ map_of_le cond h end map -lemma eq_of_proj_eq [t2_space X] [compact_space X] [disc : totally_disconnected_space X] - {x y : X} : (∀ Q : discrete_quotient X, Q.proj x = Q.proj y) → x = y := +lemma eq_of_forall_proj_eq [t2_space X] [compact_space X] [disc : totally_disconnected_space X] + {x y : X} (h : ∀ Q : discrete_quotient X, Q.proj x = Q.proj y) : x = y := begin - intro h, - change x ∈ ({y} : set X), - rw totally_disconnected_space_iff_connected_component_singleton at disc, - rw [← disc y, connected_component_eq_Inter_clopen], - rintros U ⟨⟨U, hU1, hU2⟩, rfl⟩, - replace h : _ ∨ _ := quotient.exact' (h (of_clopen hU1)), - tauto, + rw [← mem_singleton_iff, ← connected_component_eq_singleton, connected_component_eq_Inter_clopen, + mem_Inter], + rintro ⟨U, hU1, hU2⟩, + exact (quotient.exact' (h (of_clopen hU1))).mpr hU2 end -lemma fiber_le_of_le {A B : discrete_quotient X} (h : A ≤ B) (a : A) : - A.proj ⁻¹' {a} ≤ B.proj ⁻¹' {of_le h a} := +lemma fiber_subset_of_le {A B : discrete_quotient X} (h : A ≤ B) (a : A) : + A.proj ⁻¹' {a} ⊆ B.proj ⁻¹' {of_le h a} := begin - induction a, - erw [fiber_eq, fiber_eq], - tidy, + rcases A.proj_surjective a with ⟨a, rfl⟩, + rw [fiber_eq, of_le_proj, fiber_eq], + exact λ _ h', h h' end lemma exists_of_compat [compact_space X] (Qs : Π (Q : discrete_quotient X), Q) (compat : ∀ (A B : discrete_quotient X) (h : A ≤ B), of_le h (Qs _) = Qs _) : ∃ x : X, ∀ Q : discrete_quotient X, Q.proj x = Qs _ := begin - obtain ⟨x,hx⟩ := is_compact.nonempty_Inter_of_directed_nonempty_compact_closed - (λ (Q : discrete_quotient X), Q.proj ⁻¹' {Qs _}) (λ A B, _) (λ i, _) - (λ i, (fiber_closed _ _).is_compact) (λ i, fiber_closed _ _), + obtain ⟨x,hx⟩ : (⋂ Q, proj Q ⁻¹' {Qs Q}).nonempty := + is_compact.nonempty_Inter_of_directed_nonempty_compact_closed + (λ (Q : discrete_quotient X), Q.proj ⁻¹' {Qs _}) (directed_of_inf $ λ A B h, _) + (λ Q, (singleton_nonempty _).preimage Q.proj_surjective) + (λ i, (is_closed_preimage _ _).is_compact) (λ i, is_closed_preimage _ _), { refine ⟨x, λ Q, _⟩, exact hx _ ⟨Q,rfl⟩ }, - { refine ⟨A ⊓ B, λ a ha, _, λ a ha, _⟩, - { dsimp only, - erw ← compat (A ⊓ B) A inf_le_left, - exact fiber_le_of_le _ _ ha }, - { dsimp only, - erw ← compat (A ⊓ B) B inf_le_right, - exact fiber_le_of_le _ _ ha } }, - { obtain ⟨x,hx⟩ := i.proj_surjective (Qs i), - refine ⟨x,_⟩, - dsimp only, - rw [← hx, fiber_eq], - apply i.refl }, + { rw [← compat _ _ h], + exact fiber_subset_of_le _ _ }, end -noncomputable instance [compact_space X] : fintype S := +instance [compact_space X] : finite S := begin - have cond : is_compact (⊤ : set X) := is_compact_univ, - rw is_compact_iff_finite_subcover at cond, - have h := @cond S (λ s, S.proj ⁻¹' {s}) (λ s, fiber_open _ _) - (λ x hx, ⟨S.proj ⁻¹' {S.proj x}, ⟨S.proj x, rfl⟩, rfl⟩), - let T := classical.some h, - have hT := classical.some_spec h, - refine ⟨T,λ s, _⟩, - rcases S.proj_surjective s with ⟨x,rfl⟩, - rcases hT (by tauto : x ∈ ⊤) with ⟨j, ⟨j,rfl⟩, h1, ⟨hj, rfl⟩, h2⟩, - dsimp only at h2, - suffices : S.proj x = j, by rwa this, - rcases j with ⟨j⟩, - apply quotient.sound', - erw fiber_eq at h2, - exact S.symm _ _ h2 + have : compact_space S := quotient.compact_space, + rwa [← is_compact_univ_iff, is_compact_iff_finite, finite_univ_iff] at this end end discrete_quotient namespace locally_constant -variables {X} {α : Type*} (f : locally_constant X α) +variables {X} (f : locally_constant X α) /-- Any locally constant function induces a discrete quotient. -/ def discrete_quotient : discrete_quotient X := -{ rel := λ a b, f b = f a, - equiv := ⟨by tauto, by tauto, λ a b c h1 h2, by rw [h2, h1]⟩, - clopen := λ x, f.is_locally_constant.is_clopen_fiber _ } +{ to_setoid := setoid.comap f ⊥, + is_open_set_of_rel := λ x, f.is_locally_constant _ } /-- The (locally constant) function from the discrete quotient associated to a locally constant function. -/ def lift : locally_constant f.discrete_quotient α := -⟨λ a, quotient.lift_on' a f (λ a b h, h.symm), λ A, is_open_discrete _⟩ +⟨λ a, quotient.lift_on' a f (λ a b, id), λ A, is_open_discrete _⟩ @[simp] lemma lift_comp_proj : f.lift ∘ f.discrete_quotient.proj = f := by { ext, refl } From cd4242344b9edd8d646ab864292489bb2e9b81af Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 8 Feb 2023 20:27:31 +0000 Subject: [PATCH 0444/1291] feat(topology/algebra/infinite_sum): add tsum_finset_bUnion_disjoint (#18350) This is the n-ary version of `tsum_union_disjoint`, although unfortunately I realized that what I actually wanted was the n-ary version of `has_sum.add_is_compl` which is harder to state! Either way, this lemma seems worth having. --- src/topology/algebra/infinite_sum.lean | 16 ++++++++++++++++ 1 file changed, 16 insertions(+) diff --git a/src/topology/algebra/infinite_sum.lean b/src/topology/algebra/infinite_sum.lean index 069f4e8498ecd..f2a6d9e75e66d 100644 --- a/src/topology/algebra/infinite_sum.lean +++ b/src/topology/algebra/infinite_sum.lean @@ -331,6 +331,16 @@ begin exact ha.add hb end +lemma has_sum_sum_disjoint {ι} (s : finset ι) {t : ι → set β} {a : ι → α} + (hs : (s : set ι).pairwise (disjoint on t)) + (hf : ∀ i ∈ s, has_sum (f ∘ coe : t i → α) (a i)) : + has_sum (f ∘ coe : (⋃ i ∈ s, t i) → α) (∑ i in s, a i) := +begin + simp_rw has_sum_subtype_iff_indicator at *, + rw set.indicator_finset_bUnion _ _ hs, + exact has_sum_sum hf, +end + lemma has_sum.add_is_compl {s t : set β} (hs : is_compl s t) (ha : has_sum (f ∘ coe : s → α) a) (hb : has_sum (f ∘ coe : t → α) b) : has_sum f (a + b) := @@ -741,6 +751,12 @@ lemma tsum_union_disjoint {s t : set β} (hd : disjoint s t) (∑' x : s ∪ t, f x) = (∑' x : s, f x) + (∑' x : t, f x) := (hs.has_sum.add_disjoint hd ht.has_sum).tsum_eq +lemma tsum_finset_bUnion_disjoint {ι} {s : finset ι} {t : ι → set β} + (hd : (s : set ι).pairwise (disjoint on t)) + (hf : ∀ i ∈ s, summable (f ∘ coe : t i → α)) : + (∑' x : (⋃ i ∈ s, t i), f x) = ∑ i in s, ∑' x : t i, f x := +(has_sum_sum_disjoint _ hd (λ i hi, (hf i hi).has_sum)).tsum_eq + lemma tsum_even_add_odd {f : ℕ → α} (he : summable (λ k, f (2 * k))) (ho : summable (λ k, f (2 * k + 1))) : (∑' k, f (2 * k)) + (∑' k, f (2 * k + 1)) = ∑' k, f k := From c8fc41b24de310b28367a3b069bfbe4757c807b0 Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Wed, 8 Feb 2023 20:27:32 +0000 Subject: [PATCH 0445/1291] feat (analysis/special_functions): add the Beta function (#18373) Add the Beta function, and prove its relation to the Gamma function via convolution integrals. --- src/analysis/special_functions/gamma.lean | 213 +++++++++++++++++++++- 1 file changed, 210 insertions(+), 3 deletions(-) diff --git a/src/analysis/special_functions/gamma.lean b/src/analysis/special_functions/gamma.lean index 58018959827f1..3bc21785cf2eb 100644 --- a/src/analysis/special_functions/gamma.lean +++ b/src/analysis/special_functions/gamma.lean @@ -6,9 +6,10 @@ Authors: David Loeffler import measure_theory.integral.exp_decay import analysis.calculus.parametric_integral import analysis.special_functions.integrals +import analysis.convolution /-! -# The Gamma function +# The Gamma and Beta functions This file defines the `Γ` function (of a real or complex variable `s`). We define this by Euler's integral `Γ(s) = ∫ x in Ioi 0, exp (-x) * x ^ (s - 1)` in the range where this integral converges @@ -19,7 +20,7 @@ We show that this integral satisfies `Γ(1) = 1` and `Γ(s + 1) = s * Γ(s)`; he integral in the convergence range. In the complex case we also prove that the resulting function is holomorphic on `ℂ` away from the points `{-n : n ∈ ℕ}`. -## Main statements (real case) +## Gamma function: main statements (real case) * `real.Gamma` : the `Γ` function (of a real variable). * `real.Gamma_eq_integral` : for `0 < s`, `Γ(s)` agrees with Euler's integral @@ -36,7 +37,7 @@ holomorphic on `ℂ` away from the points `{-n : n ∈ ℕ}`. the unique log-convex, positive-valued function on `Ioi 0` satisfying the functional equation and having `Γ 1 = 1`. -## Main statements (complex case) +## Gamma function: main statements (complex case) * `complex.Gamma` : the `Γ` function (of a complex variable). * `complex.Gamma_eq_integral` : for `0 < re s`, `Γ(s)` agrees with Euler's integral. @@ -45,6 +46,13 @@ holomorphic on `ℂ` away from the points `{-n : n ∈ ℕ}`. * `complex.differentiable_at_Gamma` : `Γ` is complex-differentiable at all `s : ℂ` with `s ∉ {-n : n ∈ ℕ}`. +## Beta function + +* `complex.beta_integral`: the Beta function `Β(u, v)`, where `u`, `v` are complex with positive + real part. +* `complex.Gamma_mul_Gamma_eq_beta_integral`: the formula + `Gamma u * Gamma v = Gamma (u + v) * beta_integral u v`. + ## Tags Gamma @@ -1007,3 +1015,202 @@ end end strict_mono end real + +section beta_integral + +namespace complex + +notation `cexp` := complex.exp + +/-- The Beta function `Β (u, v)`, defined as `∫ x:ℝ in 0..1, x ^ (u - 1) * (1 - x) ^ (v - 1)`. -/ +noncomputable def beta_integral (u v : ℂ) : ℂ := +∫ (x:ℝ) in 0..1, x ^ (u - 1) * (1 - x) ^ (v - 1) + +/-- Auxiliary lemma for `beta_integral_convergent`, showing convergence at the left endpoint. -/ +lemma beta_integral_convergent_left {u : ℂ} (hu : 0 < re u) (v : ℂ) : + interval_integrable (λ x, x ^ (u - 1) * (1 - x) ^ (v - 1) : ℝ → ℂ) volume 0 (1 / 2) := +begin + apply interval_integrable.mul_continuous_on, + { refine interval_integral.interval_integrable_cpow' _, + rwa [sub_re, one_re, ←zero_sub, sub_lt_sub_iff_right] }, + { apply continuous_at.continuous_on, + intros x hx, + rw uIcc_of_le (by positivity: (0:ℝ) ≤ 1/2) at hx, + apply continuous_at.cpow, + { exact (continuous_const.sub continuous_of_real).continuous_at }, + { exact continuous_at_const }, + { rw [sub_re, one_re, of_real_re, sub_pos], + exact or.inl (hx.2.trans_lt (by norm_num : (1/2:ℝ) < 1)) } } +end + +/-- The Beta integral is convergent for all `u, v` of positive real part. -/ +lemma beta_integral_convergent {u v : ℂ} (hu : 0 < re u) (hv : 0 < re v) : + interval_integrable (λ x, x ^ (u - 1) * (1 - x) ^ (v - 1) : ℝ → ℂ) volume 0 1 := +begin + refine (beta_integral_convergent_left hu v).trans _, + rw interval_integrable.iff_comp_neg, + convert ((beta_integral_convergent_left hv u).comp_add_right 1).symm, + { ext1 x, + conv_lhs { rw mul_comm }, + congr' 2; + { push_cast, ring } }, + { norm_num }, + { norm_num } +end + +lemma beta_integral_symm (u v : ℂ) : + beta_integral v u = beta_integral u v := +begin + rw [beta_integral, beta_integral], + have := interval_integral.integral_comp_mul_add + (λ x:ℝ, (x:ℂ) ^ (u - 1) * (1 - ↑x) ^ (v - 1)) (neg_one_lt_zero.ne) 1, + rw [inv_neg, inv_one, neg_one_smul, ←interval_integral.integral_symm] at this, + convert this, + { ext1 x, rw mul_comm, congr; + { push_cast, ring } }, + { ring }, { ring } +end + +lemma beta_integral_eval_one_right {u : ℂ} (hu : 0 < re u) : + beta_integral u 1 = 1 / u := +begin + simp_rw [beta_integral, sub_self, cpow_zero, mul_one], + rw integral_cpow (or.inl _), + { rw [of_real_zero, of_real_one, one_cpow, zero_cpow, + sub_zero, sub_add_cancel], + rw sub_add_cancel, + contrapose! hu, rw [hu, zero_re] }, + { rwa [sub_re, one_re, ←sub_pos, sub_neg_eq_add, sub_add_cancel] }, +end + +lemma beta_integral_scaled (s t : ℂ) {a : ℝ} (ha : 0 < a) : + ∫ x in 0..a, (x:ℂ) ^ (s - 1) * (a - x) ^ (t - 1) = a ^ (s + t - 1) * beta_integral s t := +begin + have ha' : (a:ℂ) ≠ 0, from of_real_ne_zero.mpr ha.ne', + rw beta_integral, + have A : (a:ℂ) ^ (s + t - 1) = a * (a ^ (s - 1) * a ^ (t - 1)), + { rw [(by abel : s + t - 1 = 1 + (s - 1) + (t - 1)), + cpow_add _ _ ha', cpow_add 1 _ ha', cpow_one, mul_assoc] }, + rw [A, mul_assoc, ←interval_integral.integral_const_mul ((↑a) ^ _ * _), + ←real_smul, ←(zero_div a), ←div_self ha.ne', + ←interval_integral.integral_comp_div _ ha.ne', zero_div], + simp_rw interval_integral.integral_of_le ha.le, + refine set_integral_congr measurable_set_Ioc (λ x hx, _), + dsimp only, + rw mul_mul_mul_comm, + congr' 1, + { rw [←mul_cpow_of_real_nonneg ha.le (div_pos hx.1 ha).le, of_real_div, mul_div_cancel' _ ha'] }, + { rw [(by push_cast : (1:ℂ) - ↑(x / a) = ↑(1 - x / a)), + ←mul_cpow_of_real_nonneg ha.le (sub_nonneg.mpr $ (div_le_one ha).mpr hx.2)], + push_cast, + rw [mul_sub, mul_one, mul_div_cancel' _ ha'] } +end + +/-- Relation between Beta integral and Gamma function. -/ +lemma Gamma_mul_Gamma_eq_beta_integral {s t : ℂ} (hs : 0 < re s) (ht : 0 < re t) : + Gamma s * Gamma t = Gamma (s + t) * beta_integral s t := +begin + -- Note that we haven't proved (yet) that the Gamma function has no zeroes, so we can't formulate + -- this as a formula for the Beta function. + have conv_int := integral_pos_convolution (Gamma_integral_convergent hs) + (Gamma_integral_convergent ht) (continuous_linear_map.mul ℝ ℂ), + simp_rw continuous_linear_map.mul_apply' at conv_int, + have hst : 0 < re (s + t), + { rw add_re, exact add_pos hs ht }, + rw [Gamma_eq_integral hs, Gamma_eq_integral ht, Gamma_eq_integral hst, Gamma_integral, + Gamma_integral, Gamma_integral, ←conv_int, ←integral_mul_right (beta_integral _ _)], + refine set_integral_congr measurable_set_Ioi (λ x hx, _), + dsimp only, + rw [mul_assoc, ←beta_integral_scaled s t hx, ←interval_integral.integral_const_mul], + congr' 1 with y:1, + push_cast, + suffices : cexp (-x) = cexp (-y) * cexp (-(x - y)), + { rw this, ring }, + { rw ←complex.exp_add, congr' 1, abel }, +end + +/-- Recurrence formula for the Beta function. -/ +lemma beta_integral_recurrence {u v : ℂ} (hu : 0 < re u) (hv : 0 < re v) : + u * beta_integral u (v + 1) = v * beta_integral (u + 1) v := +begin + -- NB: If we knew `Gamma (u + v + 1) ≠ 0` this would be an easy consequence of + -- `Gamma_mul_Gamma_eq_beta_integral`; but we don't know that yet. We will prove it later, but + -- this lemma is needed in the proof. So we give a (somewhat laborious) direct argument. + let F : ℝ → ℂ := λ x, x ^ u * (1 - x) ^ v, + have hu' : 0 < re (u + 1), by { rw [add_re, one_re], positivity }, + have hv' : 0 < re (v + 1), by { rw [add_re, one_re], positivity }, + have hc : continuous_on F (Icc 0 1), + { refine (continuous_at.continuous_on (λ x hx, _)).mul (continuous_at.continuous_on (λ x hx, _)), + { refine (continuous_at_cpow_const_of_re_pos (or.inl _) hu).comp + continuous_of_real.continuous_at, + rw of_real_re, exact hx.1 }, + { refine (continuous_at_cpow_const_of_re_pos (or.inl _) hv).comp + (continuous_const.sub continuous_of_real).continuous_at, + rw [sub_re, one_re, of_real_re, sub_nonneg], + exact hx.2 } }, + have hder : ∀ (x : ℝ), x ∈ Ioo (0:ℝ) 1 → has_deriv_at F + (u * (↑x ^ (u - 1) * (1 - ↑x) ^ v) - v * (↑x ^ u * (1 - ↑x) ^ (v - 1))) x, + { intros x hx, + have U : has_deriv_at (λ y:ℂ, y ^ u) (u * ↑x ^ (u - 1)) ↑x, + { have := has_deriv_at.cpow_const (has_deriv_at_id ↑x) (or.inl _), + { rw mul_one at this, exact this }, + { rw [id.def, of_real_re], exact hx.1 } }, + have V : has_deriv_at (λ y:ℂ, (1 - y) ^ v) (-v * (1 - ↑x) ^ (v - 1)) ↑x, + { have A := has_deriv_at.cpow_const (has_deriv_at_id (1 - ↑x)) (or.inl _), + rotate, { exact v }, + { rw [id.def, sub_re, one_re, of_real_re, sub_pos], exact hx.2 }, + simp_rw [id.def] at A, + have B : has_deriv_at (λ y:ℂ, 1 - y) (-1) ↑x, + { apply has_deriv_at.const_sub, apply has_deriv_at_id }, + convert has_deriv_at.comp ↑x A B using 1, + ring }, + convert (U.mul V).comp_of_real, + ring }, + have h_int := ((beta_integral_convergent hu hv').const_mul u).sub + ((beta_integral_convergent hu' hv).const_mul v), + dsimp only at h_int, + rw [add_sub_cancel, add_sub_cancel] at h_int, + have int_ev := interval_integral.integral_eq_sub_of_has_deriv_at_of_le zero_le_one hc hder h_int, + have hF0 : F 0 = 0, + { simp only [mul_eq_zero, of_real_zero, cpow_eq_zero_iff, eq_self_iff_true, + ne.def, true_and, sub_zero, one_cpow, one_ne_zero, or_false], + contrapose! hu, rw [hu, zero_re] }, + have hF1 : F 1 = 0, + { simp only [mul_eq_zero, of_real_one, one_cpow, one_ne_zero, sub_self, + cpow_eq_zero_iff, eq_self_iff_true, ne.def, true_and, false_or], + contrapose! hv, rw [hv, zero_re] }, + rw [hF0, hF1, sub_zero, interval_integral.integral_sub, + interval_integral.integral_const_mul, interval_integral.integral_const_mul] at int_ev, + { rw [beta_integral, beta_integral, ←sub_eq_zero], + convert int_ev; + { ext1 x, congr, abel } }, + { apply interval_integrable.const_mul, + convert beta_integral_convergent hu hv', + ext1 x, rw add_sub_cancel }, + { apply interval_integrable.const_mul, + convert beta_integral_convergent hu' hv, + ext1 x, rw add_sub_cancel }, +end + +/-- Explicit formula for the Beta function when second argument is a positive integer. -/ +lemma beta_integral_eval_nat_add_one_right {u : ℂ} (hu : 0 < re u) (n : ℕ) : + beta_integral u (n + 1) = n! / ∏ (j:ℕ) in finset.range (n + 1), (u + j) := +begin + induction n with n IH generalizing u, + { rw [nat.cast_zero, zero_add, beta_integral_eval_one_right hu, + nat.factorial_zero, nat.cast_one, zero_add, finset.prod_range_one, nat.cast_zero, add_zero] }, + { have := beta_integral_recurrence hu (_ : 0 < re n.succ), + swap, { rw [←of_real_nat_cast, of_real_re], positivity }, + rw [mul_comm u _, ←eq_div_iff] at this, + swap, { contrapose! hu, rw [hu, zero_re] }, + rw [this, finset.prod_range_succ', nat.cast_succ, IH], + swap, { rw [add_re, one_re], positivity }, + rw [nat.factorial_succ, nat.cast_mul, nat.cast_add, nat.cast_one, nat.cast_zero, add_zero, + ←mul_div_assoc, ←div_div], + congr' 3 with j:1, + push_cast, abel } +end + +end complex + +end beta_integral From 38d38e090c9bcfba03ecbe350534d1b1a2660bd6 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 8 Feb 2023 20:27:35 +0000 Subject: [PATCH 0446/1291] chore(analysis/normed_space/pi_Lp): pi_Lp.equiv is continuous (#18402) The main benefit here is that this should now work with the continuity tactic. --- src/analysis/normed_space/pi_Lp.lean | 16 ++++++++++++++++ 1 file changed, 16 insertions(+) diff --git a/src/analysis/normed_space/pi_Lp.lean b/src/analysis/normed_space/pi_Lp.lean index 6c501a292c752..877ee48cc06c2 100644 --- a/src/analysis/normed_space/pi_Lp.lean +++ b/src/analysis/normed_space/pi_Lp.lean @@ -393,6 +393,22 @@ end aux instance uniform_space [Π i, uniform_space (β i)] : uniform_space (pi_Lp p β) := Pi.uniform_space _ +lemma uniform_continuous_equiv [Π i, uniform_space (β i)] : + uniform_continuous (pi_Lp.equiv p β) := +uniform_continuous_id + +lemma uniform_continuous_equiv_symm [Π i, uniform_space (β i)] : + uniform_continuous (pi_Lp.equiv p β).symm := +uniform_continuous_id + +@[continuity] +lemma continuous_equiv [Π i, uniform_space (β i)] : continuous (pi_Lp.equiv p β) := +continuous_id + +@[continuity] +lemma continuous_equiv_symm [Π i, uniform_space (β i)] : continuous (pi_Lp.equiv p β).symm := +continuous_id + variable [fintype ι] instance bornology [Π i, bornology (β i)] : bornology (pi_Lp p β) := pi.bornology From b6da1a0b3e7cd83b1f744c49ce48ef8c6307d2f6 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Thu, 9 Feb 2023 01:31:16 +0000 Subject: [PATCH 0447/1291] chore(*): add mathlib4 synchronization comments (#18406) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `combinatorics.simple_graph.regularity.equitabilise` * `data.dfinsupp.interval` * `data.finsupp.interval` * `data.set.mul_antidiagonal` * `group_theory.divisible` * `linear_algebra.affine_space.basic` * `order.upper_lower.hom` * `ring_theory.ore_localization.basic` * `topology.algebra.order.liminf_limsup` * `topology.filter` * `topology.order.hom.basic` --- src/combinatorics/simple_graph/regularity/equitabilise.lean | 3 +++ src/data/dfinsupp/interval.lean | 3 +++ src/data/finsupp/interval.lean | 3 +++ src/data/set/mul_antidiagonal.lean | 5 ++++- src/group_theory/divisible.lean | 3 +++ src/linear_algebra/affine_space/basic.lean | 3 +++ src/order/upper_lower/hom.lean | 3 +++ src/ring_theory/ore_localization/basic.lean | 3 +++ src/topology/algebra/order/liminf_limsup.lean | 3 +++ src/topology/filter.lean | 3 +++ src/topology/order/hom/basic.lean | 3 +++ 11 files changed, 34 insertions(+), 1 deletion(-) diff --git a/src/combinatorics/simple_graph/regularity/equitabilise.lean b/src/combinatorics/simple_graph/regularity/equitabilise.lean index 2f264fd6908dc..32e964da201a5 100644 --- a/src/combinatorics/simple_graph/regularity/equitabilise.lean +++ b/src/combinatorics/simple_graph/regularity/equitabilise.lean @@ -8,6 +8,9 @@ import order.partition.equipartition /-! # Equitabilising a partition +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file allows to blow partitions up into parts of controlled size. Given a partition `P` and `a b m : ℕ`, we want to find a partition `Q` with `a` parts of size `m` and `b` parts of size `m + 1` such that all parts of `P` are "as close as possible" to unions of parts of `Q`. By diff --git a/src/data/dfinsupp/interval.lean b/src/data/dfinsupp/interval.lean index 6f27107710034..e010360860147 100644 --- a/src/data/dfinsupp/interval.lean +++ b/src/data/dfinsupp/interval.lean @@ -11,6 +11,9 @@ import data.dfinsupp.order /-! # Finite intervals of finitely supported functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file provides the `locally_finite_order` instance for `Π₀ i, α i` when `α` itself is locally finite and calculates the cardinality of its finite intervals. -/ diff --git a/src/data/finsupp/interval.lean b/src/data/finsupp/interval.lean index 985a3572b6e3a..b71f5d4d4987c 100644 --- a/src/data/finsupp/interval.lean +++ b/src/data/finsupp/interval.lean @@ -10,6 +10,9 @@ import data.finsupp.order /-! # Finite intervals of finitely supported functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file provides the `locally_finite_order` instance for `ι →₀ α` when `α` itself is locally finite and calculates the cardinality of its finite intervals. diff --git a/src/data/set/mul_antidiagonal.lean b/src/data/set/mul_antidiagonal.lean index 1410fd8e62481..49f259ebfcde6 100644 --- a/src/data/set/mul_antidiagonal.lean +++ b/src/data/set/mul_antidiagonal.lean @@ -5,7 +5,10 @@ Authors: Johan Commelin, Floris van Doorn -/ import order.well_founded_set -/-! # Multiplication antidiagonal -/ +/-! # Multiplication antidiagonal + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4.-/ namespace set variables {α : Type*} diff --git a/src/group_theory/divisible.lean b/src/group_theory/divisible.lean index d4678db9e4375..d31c71515a3e6 100644 --- a/src/group_theory/divisible.lean +++ b/src/group_theory/divisible.lean @@ -10,6 +10,9 @@ import algebra.group.pi /-! # Divisible Group and rootable group +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we define a divisible add monoid and a rootable monoid with some basic properties. ## Main definition diff --git a/src/linear_algebra/affine_space/basic.lean b/src/linear_algebra/affine_space/basic.lean index 9df96f95457ad..091f47b8114c5 100644 --- a/src/linear_algebra/affine_space/basic.lean +++ b/src/linear_algebra/affine_space/basic.lean @@ -8,6 +8,9 @@ import algebra.add_torsor /-! # Affine space +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we introduce the following notation: * `affine_space V P` is an alternative notation for `add_torsor V P` introduced at the end of this diff --git a/src/order/upper_lower/hom.lean b/src/order/upper_lower/hom.lean index 285f0f407b18c..99096745ac062 100644 --- a/src/order/upper_lower/hom.lean +++ b/src/order/upper_lower/hom.lean @@ -9,6 +9,9 @@ import order.hom.complete_lattice /-! # `upper_set.Ici` etc as `sup`/`Sup`/`inf`/`Inf`-homomorphisms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define `upper_set.Ici_sup_hom` etc. These functions are `upper_set.Ici` and `lower_set.Iic` bundled as `sup_hom`s, `inf_hom`s, `Sup_hom`s, or `Inf_hom`s. -/ diff --git a/src/ring_theory/ore_localization/basic.lean b/src/ring_theory/ore_localization/basic.lean index 86a19adfaea4a..d307a9e61ee95 100644 --- a/src/ring_theory/ore_localization/basic.lean +++ b/src/ring_theory/ore_localization/basic.lean @@ -12,6 +12,9 @@ import tactic.noncomm_ring # Localization over right Ore sets. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the localization of a monoid over a right Ore set and proves its universal mapping property. It then extends the definition and its properties first to semirings and then to rings. We show that in the case of a commutative monoid this definition coincides with the diff --git a/src/topology/algebra/order/liminf_limsup.lean b/src/topology/algebra/order/liminf_limsup.lean index 98c5b37d1a312..8d1cbcb99ac54 100644 --- a/src/topology/algebra/order/liminf_limsup.lean +++ b/src/topology/algebra/order/liminf_limsup.lean @@ -12,6 +12,9 @@ import topology.order.basic /-! # Lemmas about liminf and limsup in an order topology. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ open filter diff --git a/src/topology/filter.lean b/src/topology/filter.lean index fa87fc9e8dd6e..2adb2b03378c2 100644 --- a/src/topology/filter.lean +++ b/src/topology/filter.lean @@ -10,6 +10,9 @@ import data.set.intervals.monotone /-! # Topology on the set of filters on a type +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file introduce topology on `filter α`. It is generated by the sets `set.Iic (𝓟 s) = {l : filter α | s ∈ l}`, `s : set α`. A set `s : set (filter α)` is open if and only if it is a union of a family of these basic open sets, see `filter.is_open_iff`. diff --git a/src/topology/order/hom/basic.lean b/src/topology/order/hom/basic.lean index a6430e9a5bdb4..b1af0090e7247 100644 --- a/src/topology/order/hom/basic.lean +++ b/src/topology/order/hom/basic.lean @@ -9,6 +9,9 @@ import topology.continuous_function.basic /-! # Continuous order homomorphisms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines continuous order homomorphisms, that is maps which are both continuous and monotone. They are also called Priestley homomorphisms because they are the morphisms of the category of Priestley spaces. From 639b66d9330da2cd8dd02be1a51b6d10d058e5e4 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 9 Feb 2023 08:59:41 +0000 Subject: [PATCH 0448/1291] chore(topology/algebra/infinite_sum): correct argument explicitness (#18403) It is quite annoying to use the current `const_smul` and `smul_const`, because Lean often can't even infer the type of this implicit argument. There was one existing place in mathlib where this became really painful (the `@` in the diff), and there are a couple more in one of my branches. --- src/analysis/analytic/isolated_zeros.lean | 4 ++-- .../inner_product_space/l2_space.lean | 2 +- .../integral/circle_integral.lean | 6 +++--- .../integral/interval_integral.lean | 2 +- .../measure/vector_measure.lean | 2 +- src/topology/algebra/infinite_sum.lean | 20 +++++++++---------- 6 files changed, 18 insertions(+), 18 deletions(-) diff --git a/src/analysis/analytic/isolated_zeros.lean b/src/analysis/analytic/isolated_zeros.lean index c311ba2b0ad6a..a46bfea457a1d 100644 --- a/src/analysis/analytic/isolated_zeros.lean +++ b/src/analysis/analytic/isolated_zeros.lean @@ -58,7 +58,7 @@ begin from finset.sum_eq_zero (λ k hk, by simp [ha k (finset.mem_range.mp hk)]), have h2 : has_sum (λ m, z ^ (m + n) • a (m + n)) s, by simpa [h1] using (has_sum_nat_add_iff' n).mpr hs, - convert @has_sum.const_smul E ℕ 𝕜 _ _ _ _ _ _ _ (z⁻¹ ^ n) h2, + convert h2.const_smul (z⁻¹ ^ n), { field_simp [pow_add, smul_smul] }, { simp only [inv_pow] } } end @@ -79,7 +79,7 @@ begin { have hxx : ∀ (n : ℕ), x⁻¹ * x ^ (n + 1) = x ^ n := λ n, by field_simp [h, pow_succ'], suffices : has_sum (λ n, x⁻¹ • x ^ (n + 1) • p.coeff (n + 1)) (x⁻¹ • (f (z₀ + x) - f z₀)), { simpa [dslope, slope, h, smul_smul, hxx] using this }, - { simpa [hp0] using ((has_sum_nat_add_iff' 1).mpr hx).const_smul } } + { simpa [hp0] using ((has_sum_nat_add_iff' 1).mpr hx).const_smul x⁻¹ } } end lemma has_fpower_series_iterate_dslope_fslope (n : ℕ) (hp : has_fpower_series_at f p z₀) : diff --git a/src/analysis/inner_product_space/l2_space.lean b/src/analysis/inner_product_space/l2_space.lean index 35b7782f0ab6d..8be2b758acb8e 100644 --- a/src/analysis/inner_product_space/l2_space.lean +++ b/src/analysis/inner_product_space/l2_space.lean @@ -186,7 +186,7 @@ protected def linear_isometry : lp G 2 →ₗᵢ[𝕜] E := map_add' := λ f g, by simp only [tsum_add (hV.summable_of_lp f) (hV.summable_of_lp g), lp.coe_fn_add, pi.add_apply, linear_isometry.map_add], map_smul' := λ c f, by simpa only [linear_isometry.map_smul, pi.smul_apply, lp.coe_fn_smul] - using tsum_const_smul (hV.summable_of_lp f), + using tsum_const_smul c (hV.summable_of_lp f), norm_map' := λ f, begin classical, -- needed for lattice instance on `finset ι`, for `filter.at_top_ne_bot` have H : 0 < (2:ℝ≥0∞).to_real := by norm_num, diff --git a/src/measure_theory/integral/circle_integral.lean b/src/measure_theory/integral/circle_integral.lean index 55b871ba566d9..9ae1dcff76095 100644 --- a/src/measure_theory/integral/circle_integral.lean +++ b/src/measure_theory/integral/circle_integral.lean @@ -530,9 +530,9 @@ begin { exact eventually_of_forall (λ _ _, (summable_geometric_of_lt_1 hwR.1 hwR.2).mul_left _) }, { simpa only [tsum_mul_left, tsum_geometric_of_lt_1 hwR.1 hwR.2] using hf.norm.mul_continuous_on continuous_on_const }, - { refine eventually_of_forall (λ θ hθ, has_sum.const_smul _), + { refine eventually_of_forall (λ θ hθ, has_sum.const_smul _ _), simp only [smul_smul], - refine has_sum.smul_const _, + refine has_sum.smul_const _ _, have : ‖w / (circle_map c R θ - c)‖ < 1, by simpa [abs_of_pos hR] using hwR.2, convert (has_sum_geometric_of_norm_lt_1 this).mul_right _, simp [← sub_sub, ← mul_inv, sub_mul, div_mul_cancel _ (circle_map_ne_center hR.ne')] } @@ -547,7 +547,7 @@ lemma has_sum_cauchy_power_series_integral {f : ℂ → E} {c : ℂ} {R : ℝ} { ((2 * π * I : ℂ)⁻¹ • ∮ z in C(c, R), (z - (c + w))⁻¹ • f z) := begin simp only [cauchy_power_series_apply], - exact (has_sum_two_pi_I_cauchy_power_series_integral hf hw).const_smul + exact (has_sum_two_pi_I_cauchy_power_series_integral hf hw).const_smul _ end /-- For any circle integrable function `f`, the power series `cauchy_power_series f c R`, `R > 0`, diff --git a/src/measure_theory/integral/interval_integral.lean b/src/measure_theory/integral/interval_integral.lean index 95fd16215a161..baece1a595a8e 100644 --- a/src/measure_theory/integral/interval_integral.lean +++ b/src/measure_theory/integral/interval_integral.lean @@ -1061,7 +1061,7 @@ begin simp only [interval_integrable_iff, interval_integral_eq_integral_uIoc, ← ae_restrict_iff' measurable_set_uIoc] at *, exact (has_sum_integral_of_dominated_convergence bound hF_meas h_bound bound_summable - bound_integrable h_lim).const_smul + bound_integrable h_lim).const_smul _, end open topological_space diff --git a/src/measure_theory/measure/vector_measure.lean b/src/measure_theory/measure/vector_measure.lean index ea7e57ae0726c..3bce12047f068 100644 --- a/src/measure_theory/measure/vector_measure.lean +++ b/src/measure_theory/measure/vector_measure.lean @@ -256,7 +256,7 @@ def smul (r : R) (v : vector_measure α M) : vector_measure α M := { measure_of' := r • v, empty' := by rw [pi.smul_apply, empty, smul_zero], not_measurable' := λ _ hi, by rw [pi.smul_apply, v.not_measurable hi, smul_zero], - m_Union' := λ _ hf₁ hf₂, has_sum.const_smul (v.m_Union hf₁ hf₂) } + m_Union' := λ _ hf₁ hf₂, has_sum.const_smul _ (v.m_Union hf₁ hf₂) } instance : has_smul R (vector_measure α M) := ⟨smul⟩ diff --git a/src/topology/algebra/infinite_sum.lean b/src/topology/algebra/infinite_sum.lean index f2a6d9e75e66d..04c92573cb782 100644 --- a/src/topology/algebra/infinite_sum.lean +++ b/src/topology/algebra/infinite_sum.lean @@ -1096,14 +1096,14 @@ variables {R : Type*} [distrib_mul_action R α] [has_continuous_const_smul R α] {f : β → α} -lemma has_sum.const_smul {a : α} {r : R} (hf : has_sum f a) : has_sum (λ z, r • f z) (r • a) := +lemma has_sum.const_smul {a : α} (r : R) (hf : has_sum f a) : has_sum (λ z, r • f z) (r • a) := hf.map (distrib_mul_action.to_add_monoid_hom α r) (continuous_const_smul r) -lemma summable.const_smul {r : R} (hf : summable f) : summable (λ z, r • f z) := -hf.has_sum.const_smul.summable +lemma summable.const_smul (r : R) (hf : summable f) : summable (λ z, r • f z) := +(hf.has_sum.const_smul r).summable -lemma tsum_const_smul [t2_space α] {r : R} (hf : summable f) : ∑' z, r • f z = r • ∑' z, f z := -hf.has_sum.const_smul.tsum_eq +lemma tsum_const_smul [t2_space α] (r : R) (hf : summable f) : ∑' z, r • f z = r • ∑' z, f z := +(hf.has_sum.const_smul r).tsum_eq end const_smul @@ -1114,14 +1114,14 @@ variables {R : Type*} [module R α] [has_continuous_smul R α] {f : β → R} -lemma has_sum.smul_const {a : α} {r : R} (hf : has_sum f r) : has_sum (λ z, f z • a) (r • a) := +lemma has_sum.smul_const {r : R} (hf : has_sum f r) (a : α) : has_sum (λ z, f z • a) (r • a) := hf.map ((smul_add_hom R α).flip a) (continuous_id.smul continuous_const) -lemma summable.smul_const {a : α} (hf : summable f) : summable (λ z, f z • a) := -hf.has_sum.smul_const.summable +lemma summable.smul_const (hf : summable f) (a : α) : summable (λ z, f z • a) := +(hf.has_sum.smul_const a).summable -lemma tsum_smul_const [t2_space α] {a : α} (hf : summable f) : ∑' z, f z • a = (∑' z, f z) • a := -hf.has_sum.smul_const.tsum_eq +lemma tsum_smul_const [t2_space α] (hf : summable f) (a : α) : ∑' z, f z • a = (∑' z, f z) • a := +(hf.has_sum.smul_const a).tsum_eq end smul_const From 916aaa1687e8c0622135f742d9165229fa8db4ea Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Thu, 9 Feb 2023 11:00:06 +0000 Subject: [PATCH 0449/1291] chore(*): remove default.lean files that had nothing but imports (#18408) --- .../100-theorems-list/82_cubing_a_cube.lean | 2 +- docs/tutorial/category_theory/Ab.lean | 2 +- .../calculating_colimits_in_Top.lean | 1 - src/algebra/big_operators/default.lean | 8 ------ src/algebra/category/Group/default.lean | 4 --- src/algebra/category/Mon/default.lean | 1 - src/algebra/category/Ring/default.lean | 3 -- src/algebra/char_p/default.lean | 5 ---- .../computation/default.lean | 14 ---------- src/algebra/continued_fractions/default.lean | 14 ---------- src/algebra/default.lean | 2 -- src/algebra/group/default.lean | 15 ---------- src/algebra/group_power/default.lean | 1 - src/algebra/group_with_zero/default.lean | 1 - src/algebra/module/default.lean | 12 -------- src/algebra/ring/default.lean | 11 -------- src/category_theory/adjunction/default.lean | 3 -- .../concrete_category/default.lean | 1 - src/category_theory/functor/default.lean | 1 - .../limits/shapes/default.lean | 12 -------- src/category_theory/monad/default.lean | 2 -- src/category_theory/products/default.lean | 2 -- src/category_theory/subobject/default.lean | 5 ---- src/category_theory/sums/default.lean | 1 - src/control/traversable/default.lean | 2 -- src/data/fin/tuple/default.lean | 1 - src/data/finite/default.lean | 1 - src/data/finset/default.lean | 10 ------- src/data/finsupp/default.lean | 10 ------- src/data/list/default.lean | 28 ------------------- src/data/multiset/default.lean | 14 ---------- src/data/mv_polynomial/default.lean | 4 --- src/data/nat/choose/default.lean | 3 -- src/data/pfunctor/univariate/default.lean | 2 -- src/data/polynomial/default.lean | 5 ---- src/data/polynomial/degree/default.lean | 1 - src/data/qpf/multivariate/default.lean | 9 ------ src/data/rat/default.lean | 9 ------ src/data/rbtree/default.lean | 6 ---- src/data/set/default.lean | 2 -- src/data/set/intervals/default.lean | 2 -- src/data/sigma/default.lean | 1 - src/field_theory/minpoly/default.lean | 3 -- src/geometry/euclidean/default.lean | 5 ---- src/group_theory/group_action/default.lean | 1 - src/group_theory/submonoid/default.lean | 1 - .../clifford_algebra/default.lean | 2 -- src/linear_algebra/default.lean | 1 - src/linear_algebra/matrix/default.lean | 8 ------ src/number_theory/padics/default.lean | 1 - src/order/default.lean | 2 -- src/order/filter/default.lean | 1 - src/ring_theory/adjoin/default.lean | 1 - src/ring_theory/polynomial/default.lean | 1 - src/topology/category/Top/default.lean | 3 -- test/delta_instance.lean | 2 +- test/free_algebra.lean | 2 +- test/library_search/ring_theory.lean | 2 +- test/localized/localized.lean | 2 +- test/mk_iff_of_inductive.lean | 1 - 60 files changed, 6 insertions(+), 266 deletions(-) delete mode 100644 src/algebra/big_operators/default.lean delete mode 100644 src/algebra/category/Group/default.lean delete mode 100644 src/algebra/category/Mon/default.lean delete mode 100644 src/algebra/category/Ring/default.lean delete mode 100644 src/algebra/char_p/default.lean delete mode 100644 src/algebra/continued_fractions/computation/default.lean delete mode 100644 src/algebra/continued_fractions/default.lean delete mode 100644 src/algebra/default.lean delete mode 100644 src/algebra/group/default.lean delete mode 100644 src/algebra/group_power/default.lean delete mode 100644 src/algebra/group_with_zero/default.lean delete mode 100644 src/algebra/module/default.lean delete mode 100644 src/algebra/ring/default.lean delete mode 100644 src/category_theory/adjunction/default.lean delete mode 100644 src/category_theory/concrete_category/default.lean delete mode 100644 src/category_theory/functor/default.lean delete mode 100644 src/category_theory/limits/shapes/default.lean delete mode 100644 src/category_theory/monad/default.lean delete mode 100644 src/category_theory/products/default.lean delete mode 100644 src/category_theory/subobject/default.lean delete mode 100644 src/category_theory/sums/default.lean delete mode 100644 src/control/traversable/default.lean delete mode 100644 src/data/fin/tuple/default.lean delete mode 100644 src/data/finite/default.lean delete mode 100644 src/data/finset/default.lean delete mode 100644 src/data/finsupp/default.lean delete mode 100644 src/data/list/default.lean delete mode 100644 src/data/multiset/default.lean delete mode 100644 src/data/mv_polynomial/default.lean delete mode 100644 src/data/nat/choose/default.lean delete mode 100644 src/data/pfunctor/univariate/default.lean delete mode 100644 src/data/polynomial/default.lean delete mode 100644 src/data/polynomial/degree/default.lean delete mode 100644 src/data/qpf/multivariate/default.lean delete mode 100644 src/data/rat/default.lean delete mode 100644 src/data/rbtree/default.lean delete mode 100644 src/data/set/default.lean delete mode 100644 src/data/set/intervals/default.lean delete mode 100644 src/data/sigma/default.lean delete mode 100644 src/field_theory/minpoly/default.lean delete mode 100644 src/geometry/euclidean/default.lean delete mode 100644 src/group_theory/group_action/default.lean delete mode 100644 src/group_theory/submonoid/default.lean delete mode 100644 src/linear_algebra/clifford_algebra/default.lean delete mode 100644 src/linear_algebra/default.lean delete mode 100644 src/linear_algebra/matrix/default.lean delete mode 100644 src/number_theory/padics/default.lean delete mode 100644 src/order/default.lean delete mode 100644 src/order/filter/default.lean delete mode 100644 src/ring_theory/adjoin/default.lean delete mode 100644 src/ring_theory/polynomial/default.lean delete mode 100644 src/topology/category/Top/default.lean diff --git a/archive/100-theorems-list/82_cubing_a_cube.lean b/archive/100-theorems-list/82_cubing_a_cube.lean index ff662f64b69eb..62061c0d53a07 100644 --- a/archive/100-theorems-list/82_cubing_a_cube.lean +++ b/archive/100-theorems-list/82_cubing_a_cube.lean @@ -5,7 +5,7 @@ Authors: Floris van Doorn -/ import data.real.basic import data.set.finite -import data.set.intervals +import data.set.intervals.disjoint /-! Proof that a cube (in dimension n ≥ 3) cannot be cubed: diff --git a/docs/tutorial/category_theory/Ab.lean b/docs/tutorial/category_theory/Ab.lean index cc6dc64ffed42..f62da617aa52d 100644 --- a/docs/tutorial/category_theory/Ab.lean +++ b/docs/tutorial/category_theory/Ab.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ -import algebra.category.Group +import algebra.category.Group.abelian import category_theory.limits.shapes.kernels noncomputable theory diff --git a/docs/tutorial/category_theory/calculating_colimits_in_Top.lean b/docs/tutorial/category_theory/calculating_colimits_in_Top.lean index d7d33ac9f1391..9adee565f3fdd 100644 --- a/docs/tutorial/category_theory/calculating_colimits_in_Top.lean +++ b/docs/tutorial/category_theory/calculating_colimits_in_Top.lean @@ -1,5 +1,4 @@ import topology.category.Top.limits -import category_theory.limits.shapes import topology.instances.real import topology.tactic diff --git a/src/algebra/big_operators/default.lean b/src/algebra/big_operators/default.lean deleted file mode 100644 index 7924b69ff6c60..0000000000000 --- a/src/algebra/big_operators/default.lean +++ /dev/null @@ -1,8 +0,0 @@ --- Import this file to pull in everything about "big operators". --- When preparing a contribution to mathlib, it is best to minimize the imports you use. -import algebra.big_operators.order -import algebra.big_operators.intervals -import algebra.big_operators.ring -import algebra.big_operators.pi -import algebra.big_operators.finsupp -import algebra.big_operators.nat_antidiagonal diff --git a/src/algebra/category/Group/default.lean b/src/algebra/category/Group/default.lean deleted file mode 100644 index 6ca65b0ef36e4..0000000000000 --- a/src/algebra/category/Group/default.lean +++ /dev/null @@ -1,4 +0,0 @@ -import algebra.category.Group.limits -import algebra.category.Group.colimits -import algebra.category.Group.preadditive -import algebra.category.Group.zero diff --git a/src/algebra/category/Mon/default.lean b/src/algebra/category/Mon/default.lean deleted file mode 100644 index dabba273e243d..0000000000000 --- a/src/algebra/category/Mon/default.lean +++ /dev/null @@ -1 +0,0 @@ -import algebra.category.Mon.colimits diff --git a/src/algebra/category/Ring/default.lean b/src/algebra/category/Ring/default.lean deleted file mode 100644 index ef2d0c7f26ad5..0000000000000 --- a/src/algebra/category/Ring/default.lean +++ /dev/null @@ -1,3 +0,0 @@ -import algebra.category.Ring.adjunctions -import algebra.category.Ring.limits -import algebra.category.Ring.colimits diff --git a/src/algebra/char_p/default.lean b/src/algebra/char_p/default.lean deleted file mode 100644 index f0898fa6bdf41..0000000000000 --- a/src/algebra/char_p/default.lean +++ /dev/null @@ -1,5 +0,0 @@ -import algebra.char_p.algebra -import algebra.char_p.basic -import algebra.char_p.pi -import algebra.char_p.quotient -import algebra.char_p.subring diff --git a/src/algebra/continued_fractions/computation/default.lean b/src/algebra/continued_fractions/computation/default.lean deleted file mode 100644 index a35851c11e618..0000000000000 --- a/src/algebra/continued_fractions/computation/default.lean +++ /dev/null @@ -1,14 +0,0 @@ -/- -Copyright (c) 2020 Kevin Kappelmann. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Kevin Kappelmann --/ -import algebra.continued_fractions.computation.basic -import algebra.continued_fractions.computation.translations -import algebra.continued_fractions.computation.correctness_terminating -import algebra.continued_fractions.computation.approximations -import algebra.continued_fractions.computation.terminates_iff_rat -import algebra.continued_fractions.computation.approximation_corollaries -/-! -# Default Exports for the Computation of Continued Fractions --/ diff --git a/src/algebra/continued_fractions/default.lean b/src/algebra/continued_fractions/default.lean deleted file mode 100644 index 74b9c92c2de58..0000000000000 --- a/src/algebra/continued_fractions/default.lean +++ /dev/null @@ -1,14 +0,0 @@ -/- -Copyright (c) 2019 Kevin Kappelmann. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Kevin Kappelmann --/ -import algebra.continued_fractions.basic -import algebra.continued_fractions.translations -import algebra.continued_fractions.continuants_recurrence -import algebra.continued_fractions.terminated_stable -import algebra.continued_fractions.convergents_equiv -import algebra.continued_fractions.computation -/-! -# Default Exports for Continued Fractions --/ diff --git a/src/algebra/default.lean b/src/algebra/default.lean deleted file mode 100644 index a3d2c358a9bfc..0000000000000 --- a/src/algebra/default.lean +++ /dev/null @@ -1,2 +0,0 @@ -import algebra.group -import algebra.module.basic diff --git a/src/algebra/group/default.lean b/src/algebra/group/default.lean deleted file mode 100644 index ec49ff874e6c0..0000000000000 --- a/src/algebra/group/default.lean +++ /dev/null @@ -1,15 +0,0 @@ -/- -Copyright (c) 2014 Jeremy Avigad. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Jeremy Avigad, Leonardo de Moura, Michael Howes --/ -import algebra.group.conj -import algebra.group.type_tags -import algebra.group.with_one.basic -import algebra.hom.units - -/-! -# Various multiplicative and additive structures. - -This file `import`s all files in this subdirectory except for `prod`. --/ diff --git a/src/algebra/group_power/default.lean b/src/algebra/group_power/default.lean deleted file mode 100644 index 0f989844882c4..0000000000000 --- a/src/algebra/group_power/default.lean +++ /dev/null @@ -1 +0,0 @@ -import algebra.group_power.lemmas diff --git a/src/algebra/group_with_zero/default.lean b/src/algebra/group_with_zero/default.lean deleted file mode 100644 index 03893e3bbd7b6..0000000000000 --- a/src/algebra/group_with_zero/default.lean +++ /dev/null @@ -1 +0,0 @@ -import algebra.group_with_zero.basic diff --git a/src/algebra/module/default.lean b/src/algebra/module/default.lean deleted file mode 100644 index ea1a9eeb45081..0000000000000 --- a/src/algebra/module/default.lean +++ /dev/null @@ -1,12 +0,0 @@ -/- -Copyright (c) 2020 Chris Hughes. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Chris Hughes --/ -import algebra.module.basic -import algebra.module.submodule.basic - -/-! -# Default file for module -This file imports `algebra.module.basic` and `algebra.module.submodule`. --/ diff --git a/src/algebra/ring/default.lean b/src/algebra/ring/default.lean deleted file mode 100644 index 9e07db37e3ae3..0000000000000 --- a/src/algebra/ring/default.lean +++ /dev/null @@ -1,11 +0,0 @@ -/- -Copyright (c) 2020 Chris Hughes. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Chris Hughes --/ -import algebra.ring.basic - -/-! -# Default file for ring -This file imports `algebra.ring.basic` --/ diff --git a/src/category_theory/adjunction/default.lean b/src/category_theory/adjunction/default.lean deleted file mode 100644 index d314185bc87f6..0000000000000 --- a/src/category_theory/adjunction/default.lean +++ /dev/null @@ -1,3 +0,0 @@ -import category_theory.adjunction.limits -import category_theory.adjunction.opposites -import category_theory.adjunction.reflective diff --git a/src/category_theory/concrete_category/default.lean b/src/category_theory/concrete_category/default.lean deleted file mode 100644 index 72e4679426489..0000000000000 --- a/src/category_theory/concrete_category/default.lean +++ /dev/null @@ -1 +0,0 @@ -import category_theory.concrete_category.unbundled_hom diff --git a/src/category_theory/functor/default.lean b/src/category_theory/functor/default.lean deleted file mode 100644 index 94fa892d0119c..0000000000000 --- a/src/category_theory/functor/default.lean +++ /dev/null @@ -1 +0,0 @@ -import category_theory.functor.basic diff --git a/src/category_theory/limits/shapes/default.lean b/src/category_theory/limits/shapes/default.lean deleted file mode 100644 index ce39844b6f1e1..0000000000000 --- a/src/category_theory/limits/shapes/default.lean +++ /dev/null @@ -1,12 +0,0 @@ -import category_theory.limits.shapes.terminal -import category_theory.limits.shapes.binary_products -import category_theory.limits.shapes.products -import category_theory.limits.shapes.finite_products -import category_theory.limits.shapes.finite_limits -import category_theory.limits.shapes.biproducts -import category_theory.limits.shapes.images -import category_theory.limits.shapes.zero_morphisms -import category_theory.limits.shapes.kernels -import category_theory.limits.shapes.equalizers -import category_theory.limits.shapes.wide_pullbacks -import category_theory.limits.shapes.pullbacks diff --git a/src/category_theory/monad/default.lean b/src/category_theory/monad/default.lean deleted file mode 100644 index 9068ea7602cc2..0000000000000 --- a/src/category_theory/monad/default.lean +++ /dev/null @@ -1,2 +0,0 @@ -import category_theory.monad.limits -import category_theory.monad.types diff --git a/src/category_theory/products/default.lean b/src/category_theory/products/default.lean deleted file mode 100644 index 816bf1996a348..0000000000000 --- a/src/category_theory/products/default.lean +++ /dev/null @@ -1,2 +0,0 @@ -import category_theory.products.bifunctor -import category_theory.products.associator diff --git a/src/category_theory/subobject/default.lean b/src/category_theory/subobject/default.lean deleted file mode 100644 index ac3af1749199b..0000000000000 --- a/src/category_theory/subobject/default.lean +++ /dev/null @@ -1,5 +0,0 @@ -import category_theory.subobject.mono_over -import category_theory.subobject.basic -import category_theory.subobject.factor_thru -import category_theory.subobject.well_powered -import category_theory.subobject.lattice diff --git a/src/category_theory/sums/default.lean b/src/category_theory/sums/default.lean deleted file mode 100644 index bc466880ac5de..0000000000000 --- a/src/category_theory/sums/default.lean +++ /dev/null @@ -1 +0,0 @@ -import category_theory.sums.associator diff --git a/src/control/traversable/default.lean b/src/control/traversable/default.lean deleted file mode 100644 index 9a41fa82e3fcc..0000000000000 --- a/src/control/traversable/default.lean +++ /dev/null @@ -1,2 +0,0 @@ -import control.traversable.instances -import control.traversable.lemmas diff --git a/src/data/fin/tuple/default.lean b/src/data/fin/tuple/default.lean deleted file mode 100644 index 8254c590aabd2..0000000000000 --- a/src/data/fin/tuple/default.lean +++ /dev/null @@ -1 +0,0 @@ -import data.fin.tuple.basic diff --git a/src/data/finite/default.lean b/src/data/finite/default.lean deleted file mode 100644 index c30836d0ce0cb..0000000000000 --- a/src/data/finite/default.lean +++ /dev/null @@ -1 +0,0 @@ -import data.finite.set diff --git a/src/data/finset/default.lean b/src/data/finset/default.lean deleted file mode 100644 index b8d3d55dabbaf..0000000000000 --- a/src/data/finset/default.lean +++ /dev/null @@ -1,10 +0,0 @@ -import data.finset.basic -import data.finset.fold -import data.finset.image -import data.finset.lattice -import data.finset.locally_finite -import data.finset.nat_antidiagonal -import data.finset.pi -import data.finset.powerset -import data.finset.sort -import data.finset.preimage diff --git a/src/data/finsupp/default.lean b/src/data/finsupp/default.lean deleted file mode 100644 index 1853d56cd28fa..0000000000000 --- a/src/data/finsupp/default.lean +++ /dev/null @@ -1,10 +0,0 @@ -/- -Copyright (c) 2020 Chris Hughes. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Chris Hughes --/ -import data.finsupp.defs -/-! - # Default finsupp file - This file imports `data.finsupp.basic` --/ diff --git a/src/data/list/default.lean b/src/data/list/default.lean deleted file mode 100644 index aeda587b9cf86..0000000000000 --- a/src/data/list/default.lean +++ /dev/null @@ -1,28 +0,0 @@ -/- -Copyright (c) 2014 Microsoft Corporation. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Jeremy Avigad --/ -import data.list.alist -import data.list.basic -import data.list.chain -import data.list.defs -import data.list.dedup -import data.list.forall2 -import data.list.func -import data.list.intervals -import data.list.lattice -import data.list.min_max -import data.list.indexes -import data.list.nat_antidiagonal -import data.list.nodup -import data.list.of_fn -import data.list.pairwise -import data.list.perm -import data.list.range -import data.list.rotate -import data.list.sections -import data.list.sigma -import data.list.sort -import data.list.tfae -import data.list.zip diff --git a/src/data/multiset/default.lean b/src/data/multiset/default.lean deleted file mode 100644 index e73adf88fc396..0000000000000 --- a/src/data/multiset/default.lean +++ /dev/null @@ -1,14 +0,0 @@ -import data.multiset.antidiagonal -import data.multiset.basic -import data.multiset.dedup -import data.multiset.finset_ops -import data.multiset.fold -import data.multiset.functor -import data.multiset.lattice -import data.multiset.locally_finite -import data.multiset.nat_antidiagonal -import data.multiset.nodup -import data.multiset.pi -import data.multiset.powerset -import data.multiset.sections -import data.multiset.sort diff --git a/src/data/mv_polynomial/default.lean b/src/data/mv_polynomial/default.lean deleted file mode 100644 index e90a087f351d3..0000000000000 --- a/src/data/mv_polynomial/default.lean +++ /dev/null @@ -1,4 +0,0 @@ -import data.mv_polynomial.basic -import data.mv_polynomial.variables -import data.mv_polynomial.rename -import data.mv_polynomial.comm_ring diff --git a/src/data/nat/choose/default.lean b/src/data/nat/choose/default.lean deleted file mode 100644 index c8003fa521f64..0000000000000 --- a/src/data/nat/choose/default.lean +++ /dev/null @@ -1,3 +0,0 @@ -import data.nat.choose.dvd -import data.nat.choose.cast -import data.nat.choose.sum diff --git a/src/data/pfunctor/univariate/default.lean b/src/data/pfunctor/univariate/default.lean deleted file mode 100644 index 157e8d4adb6cd..0000000000000 --- a/src/data/pfunctor/univariate/default.lean +++ /dev/null @@ -1,2 +0,0 @@ -import data.pfunctor.univariate.basic -import data.pfunctor.univariate.M diff --git a/src/data/polynomial/default.lean b/src/data/polynomial/default.lean deleted file mode 100644 index 05858af734dd8..0000000000000 --- a/src/data/polynomial/default.lean +++ /dev/null @@ -1,5 +0,0 @@ -import data.polynomial.algebra_map -import data.polynomial.field_division -import data.polynomial.derivative -import data.polynomial.identities -import data.polynomial.integral_normalization diff --git a/src/data/polynomial/degree/default.lean b/src/data/polynomial/degree/default.lean deleted file mode 100644 index 049cf1a50ebf9..0000000000000 --- a/src/data/polynomial/degree/default.lean +++ /dev/null @@ -1 +0,0 @@ -import data.polynomial.degree.lemmas diff --git a/src/data/qpf/multivariate/default.lean b/src/data/qpf/multivariate/default.lean deleted file mode 100644 index 60fe082964167..0000000000000 --- a/src/data/qpf/multivariate/default.lean +++ /dev/null @@ -1,9 +0,0 @@ - -import data.qpf.multivariate.basic -import data.qpf.multivariate.constructions.fix -import data.qpf.multivariate.constructions.cofix -import data.qpf.multivariate.constructions.comp -import data.qpf.multivariate.constructions.quot -import data.qpf.multivariate.constructions.prj -import data.qpf.multivariate.constructions.const -import data.qpf.multivariate.constructions.sigma diff --git a/src/data/rat/default.lean b/src/data/rat/default.lean deleted file mode 100644 index 8b8f15ad95b88..0000000000000 --- a/src/data/rat/default.lean +++ /dev/null @@ -1,9 +0,0 @@ -/- -Copyright (c) 2019 Kevin Kappelmann. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Kevin Kappelmann --/ -import data.rat.floor -/-! -# Default Imports to Work With Rational Numbers --/ diff --git a/src/data/rbtree/default.lean b/src/data/rbtree/default.lean deleted file mode 100644 index 3e4107221da67..0000000000000 --- a/src/data/rbtree/default.lean +++ /dev/null @@ -1,6 +0,0 @@ -/- -Copyright (c) 2017 Microsoft Corporation. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Leonardo de Moura --/ -import data.rbtree.main diff --git a/src/data/set/default.lean b/src/data/set/default.lean deleted file mode 100644 index ca6ef04fcd3cc..0000000000000 --- a/src/data/set/default.lean +++ /dev/null @@ -1,2 +0,0 @@ -import data.set.finite -import data.set.intervals diff --git a/src/data/set/intervals/default.lean b/src/data/set/intervals/default.lean deleted file mode 100644 index b071ccd6fd329..0000000000000 --- a/src/data/set/intervals/default.lean +++ /dev/null @@ -1,2 +0,0 @@ -import data.set.intervals.disjoint -import data.set.intervals.unordered_interval diff --git a/src/data/sigma/default.lean b/src/data/sigma/default.lean deleted file mode 100644 index 98ed296c33440..0000000000000 --- a/src/data/sigma/default.lean +++ /dev/null @@ -1 +0,0 @@ -import data.sigma.basic diff --git a/src/field_theory/minpoly/default.lean b/src/field_theory/minpoly/default.lean deleted file mode 100644 index e7bc53a8783e6..0000000000000 --- a/src/field_theory/minpoly/default.lean +++ /dev/null @@ -1,3 +0,0 @@ -import field_theory.minpoly.basic -import field_theory.minpoly.field -import field_theory.minpoly.is_integrally_closed diff --git a/src/geometry/euclidean/default.lean b/src/geometry/euclidean/default.lean deleted file mode 100644 index 9002867580ffe..0000000000000 --- a/src/geometry/euclidean/default.lean +++ /dev/null @@ -1,5 +0,0 @@ -import geometry.euclidean.angle.unoriented.right_angle -import geometry.euclidean.basic -import geometry.euclidean.circumcenter -import geometry.euclidean.monge_point -import geometry.euclidean.triangle diff --git a/src/group_theory/group_action/default.lean b/src/group_theory/group_action/default.lean deleted file mode 100644 index f64b8ea3262f6..0000000000000 --- a/src/group_theory/group_action/default.lean +++ /dev/null @@ -1 +0,0 @@ -import group_theory.group_action.basic diff --git a/src/group_theory/submonoid/default.lean b/src/group_theory/submonoid/default.lean deleted file mode 100644 index da25bc98ded11..0000000000000 --- a/src/group_theory/submonoid/default.lean +++ /dev/null @@ -1 +0,0 @@ -import group_theory.submonoid.membership diff --git a/src/linear_algebra/clifford_algebra/default.lean b/src/linear_algebra/clifford_algebra/default.lean deleted file mode 100644 index 0157342a1887e..0000000000000 --- a/src/linear_algebra/clifford_algebra/default.lean +++ /dev/null @@ -1,2 +0,0 @@ -import linear_algebra.clifford_algebra.basic -import linear_algebra.clifford_algebra.conjugation diff --git a/src/linear_algebra/default.lean b/src/linear_algebra/default.lean deleted file mode 100644 index c33b018f641d8..0000000000000 --- a/src/linear_algebra/default.lean +++ /dev/null @@ -1 +0,0 @@ -import linear_algebra.basic diff --git a/src/linear_algebra/matrix/default.lean b/src/linear_algebra/matrix/default.lean deleted file mode 100644 index 6ac169c9a9c9d..0000000000000 --- a/src/linear_algebra/matrix/default.lean +++ /dev/null @@ -1,8 +0,0 @@ -import linear_algebra.determinant -import linear_algebra.matrix.basis -import linear_algebra.matrix.block -import linear_algebra.matrix.diagonal -import linear_algebra.matrix.dot_product -import linear_algebra.matrix.to_linear_equiv -import linear_algebra.matrix.reindex -import linear_algebra.trace diff --git a/src/number_theory/padics/default.lean b/src/number_theory/padics/default.lean deleted file mode 100644 index 9eaf5cc607bb8..0000000000000 --- a/src/number_theory/padics/default.lean +++ /dev/null @@ -1 +0,0 @@ -import number_theory.padics.padic_integers diff --git a/src/order/default.lean b/src/order/default.lean deleted file mode 100644 index b565d70b2b21c..0000000000000 --- a/src/order/default.lean +++ /dev/null @@ -1,2 +0,0 @@ -import order.boolean_algebra -import order.complete_lattice diff --git a/src/order/filter/default.lean b/src/order/filter/default.lean deleted file mode 100644 index 20469ede615cc..0000000000000 --- a/src/order/filter/default.lean +++ /dev/null @@ -1 +0,0 @@ -import order.filter.partial diff --git a/src/ring_theory/adjoin/default.lean b/src/ring_theory/adjoin/default.lean deleted file mode 100644 index 775aeca5c20d8..0000000000000 --- a/src/ring_theory/adjoin/default.lean +++ /dev/null @@ -1 +0,0 @@ -import ring_theory.adjoin.basic diff --git a/src/ring_theory/polynomial/default.lean b/src/ring_theory/polynomial/default.lean deleted file mode 100644 index 0b78bd13fb19a..0000000000000 --- a/src/ring_theory/polynomial/default.lean +++ /dev/null @@ -1 +0,0 @@ -import ring_theory.polynomial.basic diff --git a/src/topology/category/Top/default.lean b/src/topology/category/Top/default.lean deleted file mode 100644 index 9e7a0f40fe13c..0000000000000 --- a/src/topology/category/Top/default.lean +++ /dev/null @@ -1,3 +0,0 @@ -import topology.category.Top.limits -import topology.category.Top.epi_mono -import topology.category.Top.open_nhds diff --git a/test/delta_instance.lean b/test/delta_instance.lean index 41110f7e75785..32a1bb7538793 100644 --- a/test/delta_instance.lean +++ b/test/delta_instance.lean @@ -3,7 +3,7 @@ Copyright (c) 2019 Robert Y. Lewis. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Robert Y. Lewis -/ -import data.set +import data.set.basic import algebra.category.Mon.basic def X : Type := set ℕ diff --git a/test/free_algebra.lean b/test/free_algebra.lean index b570e397a309c..b8bdac6c6853f 100644 --- a/test/free_algebra.lean +++ b/test/free_algebra.lean @@ -5,7 +5,7 @@ Authors: Eric Wieser -/ import linear_algebra.exterior_algebra.basic -import linear_algebra.clifford_algebra +import linear_algebra.clifford_algebra.basic /-! Tests that the ring instances for `free_algebra` and derived quotient types actually work. diff --git a/test/library_search/ring_theory.lean b/test/library_search/ring_theory.lean index 11c3534b0b6e2..b483f4a48a992 100644 --- a/test/library_search/ring_theory.lean +++ b/test/library_search/ring_theory.lean @@ -5,7 +5,7 @@ Authors: Scott Morrison -/ import tactic.suggest import ring_theory.principal_ideal_domain -import ring_theory.polynomial +import ring_theory.polynomial.basic open_locale polynomial /- Turn off trace messages so they don't pollute the test build: -/ diff --git a/test/localized/localized.lean b/test/localized/localized.lean index 07fad09ff60b7..6d7e4030ebc25 100644 --- a/test/localized/localized.lean +++ b/test/localized/localized.lean @@ -1,5 +1,5 @@ import tactic.localized -import algebra.group_power +import algebra.group_power.lemmas open tactic local infix ` ⊹ `:59 := nat.mul diff --git a/test/mk_iff_of_inductive.lean b/test/mk_iff_of_inductive.lean index 0b8b1b6ffc908..f8bafe42a7c89 100644 --- a/test/mk_iff_of_inductive.lean +++ b/test/mk_iff_of_inductive.lean @@ -1,4 +1,3 @@ -import data.list import data.list.perm import data.multiset.basic From ccf84e0d918668460a34aa19d02fe2e0e2286da0 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 9 Feb 2023 13:14:25 +0000 Subject: [PATCH 0450/1291] feat(analysis/special_functions/trigonometric/series): express `cos`/`sin` as infinite sums (#18352) I wasn't able to find a nice way to prove this with `has_sum.even_add_odd`, so instead I transported across `nat.div_mod_equiv` such that I could split even and odd terms. If nothing else, this provides a nice example of how to use similar equivs to split by remainders upon division by values other than 2. [Zulip](https://leanprover.zulipchat.com/#narrow/stream/217875-Is-there-code-for-X.3F/topic/real.2Ecos.20and.20real.2Esin.20as.20infinite.20series/near/325193524) Co-authored-by: sgouezel --- docs/undergrad.yaml | 2 +- src/analysis/normed_space/exponential.lean | 6 + .../special_functions/exponential.lean | 17 +-- .../trigonometric/series.lean | 115 ++++++++++++++++++ 4 files changed, 123 insertions(+), 17 deletions(-) create mode 100644 src/analysis/special_functions/trigonometric/series.lean diff --git a/docs/undergrad.yaml b/docs/undergrad.yaml index c984b13916c83..450ce4b9a2365 100644 --- a/docs/undergrad.yaml +++ b/docs/undergrad.yaml @@ -370,7 +370,7 @@ Single Variable Complex Analysis: antiderivative: '' complex exponential: 'complex.exp' extension of trigonometric functions to the complex plane: 'complex.sin' - power series expansion of elementary functions: '' + power series expansion of elementary functions: 'complex.has_sum_cos' Functions on one complex variable: holomorphic functions: 'differentiable_on' Cauchy-Riemann conditions: '' diff --git a/src/analysis/normed_space/exponential.lean b/src/analysis/normed_space/exponential.lean index 6e30c8039fe74..31473c07460cc 100644 --- a/src/analysis/normed_space/exponential.lean +++ b/src/analysis/normed_space/exponential.lean @@ -623,4 +623,10 @@ end lemma exp_ℝ_ℂ_eq_exp_ℂ_ℂ : (exp ℝ : ℂ → ℂ) = exp ℂ := exp_eq_exp ℝ ℂ ℂ +/-- A version of `complex.of_real_exp` for `exp` instead of `complex.exp` -/ +@[simp, norm_cast] +lemma of_real_exp_ℝ_ℝ (r : ℝ) : ↑(exp ℝ r) = exp ℂ (r : ℂ) := +(map_exp ℝ (algebra_map ℝ ℂ) (continuous_algebra_map _ _) r).trans + (congr_fun exp_ℝ_ℂ_eq_exp_ℂ_ℂ _) + end scalar_tower diff --git a/src/analysis/special_functions/exponential.lean b/src/analysis/special_functions/exponential.lean index cc3a4dcc3fc7d..d7e81c0cf58ce 100644 --- a/src/analysis/special_functions/exponential.lean +++ b/src/analysis/special_functions/exponential.lean @@ -202,8 +202,6 @@ has_strict_deriv_at_exp_zero.has_deriv_at end deriv_R_or_C -section complex - lemma complex.exp_eq_exp_ℂ : complex.exp = exp ℂ := begin refine funext (λ x, _), @@ -212,18 +210,5 @@ begin (exp_series_div_summable ℝ x).has_sum.tendsto_sum_nat end -end complex - -section real - lemma real.exp_eq_exp_ℝ : real.exp = exp ℝ := -begin - refine funext (λ x, _), - rw [real.exp, complex.exp_eq_exp_ℂ, ← exp_ℝ_ℂ_eq_exp_ℂ_ℂ, exp_eq_tsum, exp_eq_tsum_div, - ← re_to_complex, ← re_clm_apply, re_clm.map_tsum (exp_series_summable' (x : ℂ))], - refine tsum_congr (λ n, _), - rw [re_clm.map_smul, ← complex.of_real_pow, re_clm_apply, re_to_complex, complex.of_real_re, - smul_eq_mul, div_eq_inv_mul] -end - -end real +by { ext x, exact_mod_cast congr_fun complex.exp_eq_exp_ℂ x } diff --git a/src/analysis/special_functions/trigonometric/series.lean b/src/analysis/special_functions/trigonometric/series.lean new file mode 100644 index 0000000000000..0665bc3e9122f --- /dev/null +++ b/src/analysis/special_functions/trigonometric/series.lean @@ -0,0 +1,115 @@ +/- +Copyright (c) 2023 Eric Wieser. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Eric Wieser +-/ +import analysis.special_functions.exponential +/-! +# Trigonometric functions as sums of infinite series + +In this file we express trigonometric functions in terms of their series expansion. + +## Main results + +* `complex.has_sum_cos`, `complex.tsum_cos`: `complex.cos` as the sum of an infinite series. +* `real.has_sum_cos`, `real.tsum_cos`: `real.cos` as the sum of an infinite series. +* `complex.has_sum_sin`, `complex.tsum_sin`: `complex.sin` as the sum of an infinite series. +* `real.has_sum_sin`, `real.tsum_sin`: `real.sin` as the sum of an infinite series. +-/ + +open_locale nat + +/-! ### `cos` and `sin` for `ℝ` and `ℂ` -/ + +section sin_cos + +lemma complex.has_sum_cos' (z : ℂ) : + has_sum (λ n : ℕ, (z * complex.I) ^ (2 * n) / ↑(2 * n)!) (complex.cos z) := +begin + rw [complex.cos, complex.exp_eq_exp_ℂ], + have := ((exp_series_div_has_sum_exp ℂ (z * complex.I)).add + (exp_series_div_has_sum_exp ℂ (-z * complex.I))).div_const 2, + replace := ((nat.div_mod_equiv 2)).symm.has_sum_iff.mpr this, + dsimp [function.comp] at this, + simp_rw [←mul_comm 2 _] at this, + refine this.prod_fiberwise (λ k, _), + dsimp only, + convert has_sum_fintype (_ : fin 2 → ℂ) using 1, + rw fin.sum_univ_two, + simp_rw [fin.coe_zero, fin.coe_one, add_zero, pow_succ', pow_mul, + mul_pow, neg_sq, ←two_mul, neg_mul, mul_neg, neg_div, add_right_neg, zero_div, add_zero, + mul_div_cancel_left _ (two_ne_zero : (2 : ℂ) ≠ 0)], +end + +lemma complex.has_sum_sin' (z : ℂ) : + has_sum (λ n : ℕ, (z * complex.I) ^ (2 * n + 1) / ↑(2 * n + 1)! / complex.I) (complex.sin z) := +begin + rw [complex.sin, complex.exp_eq_exp_ℂ], + have := (((exp_series_div_has_sum_exp ℂ (-z * complex.I)).sub + (exp_series_div_has_sum_exp ℂ (z * complex.I))).mul_right complex.I).div_const 2, + replace := ((nat.div_mod_equiv 2)).symm.has_sum_iff.mpr this, + dsimp [function.comp] at this, + simp_rw [←mul_comm 2 _] at this, + refine this.prod_fiberwise (λ k, _), + dsimp only, + convert has_sum_fintype (_ : fin 2 → ℂ) using 1, + rw fin.sum_univ_two, + simp_rw [fin.coe_zero, fin.coe_one, add_zero, pow_succ', pow_mul, + mul_pow, neg_sq, sub_self, zero_mul, zero_div, zero_add, + neg_mul, mul_neg, neg_div, ← neg_add', ←two_mul, neg_mul, neg_div, mul_assoc, + mul_div_cancel_left _ (two_ne_zero : (2 : ℂ) ≠ 0), complex.div_I], +end + +/-- The power series expansion of `complex.cos`. -/ +lemma complex.has_sum_cos (z : ℂ) : + has_sum (λ n : ℕ, ((-1) ^ n) * z ^ (2 * n) / ↑(2 * n)!) (complex.cos z) := +begin + convert complex.has_sum_cos' z using 1, + simp_rw [mul_pow, pow_mul, complex.I_sq, mul_comm] +end + +/-- The power series expansion of `complex.sin`. -/ +lemma complex.has_sum_sin (z : ℂ) : + has_sum (λ n : ℕ, ((-1) ^ n) * z ^ (2 * n + 1) / ↑(2 * n + 1)!) (complex.sin z) := +begin + convert complex.has_sum_sin' z using 1, + simp_rw [mul_pow, pow_succ', pow_mul, complex.I_sq, ←mul_assoc, + mul_div_assoc, div_right_comm, div_self complex.I_ne_zero, mul_comm _ ((-1 : ℂ)^_), mul_one_div, + mul_div_assoc, mul_assoc] +end + +lemma complex.cos_eq_tsum' (z : ℂ) : + complex.cos z = ∑' n : ℕ, (z * complex.I) ^ (2 * n) / ↑(2 * n)! := +(complex.has_sum_cos' z).tsum_eq.symm + +lemma complex.sin_eq_tsum' (z : ℂ) : + complex.sin z = ∑' n : ℕ, (z * complex.I) ^ (2 * n + 1) / ↑(2 * n + 1)! / complex.I := +(complex.has_sum_sin' z).tsum_eq.symm + +lemma complex.cos_eq_tsum (z : ℂ) : + complex.cos z = ∑' n : ℕ, ((-1) ^ n) * z ^ (2 * n) / ↑(2 * n)! := +(complex.has_sum_cos z).tsum_eq.symm + +lemma complex.sin_eq_tsum (z : ℂ) : + complex.sin z = ∑' n : ℕ, ((-1) ^ n) * z ^ (2 * n + 1) / ↑(2 * n + 1)! := +(complex.has_sum_sin z).tsum_eq.symm + +/-- The power series expansion of `real.cos`. -/ +lemma real.has_sum_cos (r : ℝ) : + has_sum (λ n : ℕ, ((-1) ^ n) * r ^ (2 * n) / ↑(2 * n)!) (real.cos r) := +by exact_mod_cast complex.has_sum_cos r + +/-- The power series expansion of `real.sin`. -/ +lemma real.has_sum_sin (r : ℝ) : + has_sum (λ n : ℕ, ((-1) ^ n) * r ^ (2 * n + 1) / ↑(2 * n + 1)!) (real.sin r) := +by exact_mod_cast complex.has_sum_sin r + +lemma real.cos_eq_tsum (r : ℝ) : + real.cos r = ∑' n : ℕ, ((-1) ^ n) * r ^ (2 * n) / ↑(2 * n)! := +(real.has_sum_cos r).tsum_eq.symm + +lemma real.sin_eq_tsum (r : ℝ) : + real.sin r = ∑' n : ℕ, ((-1) ^ n) * r ^ (2 * n + 1) / ↑(2 * n + 1)! := +(real.has_sum_sin r).tsum_eq.symm + +end sin_cos From a75898643b2d774cced9ae7c0b28c21663b99666 Mon Sep 17 00:00:00 2001 From: sgouezel Date: Thu, 9 Feb 2023 14:33:20 +0000 Subject: [PATCH 0451/1291] feat(measure_theory/integral/set_integral): remove or weaken some integrability assumptions (#18395) When dealing with integrals on sets, we assume almost always that the set is measurable or null-measurable. However, there are situations where these assumptions can be removed, making the lemmas easier to apply, but at the cost of more involved proofs. We implement this in this PR. As an example, we prove `set_integral_eq_of_subset_of_forall_diff_eq_zero`: if `s` is contained in `t`, and `t` is null-measurable and a function vanishes on `t \ s`, then its integrals on `s` and `t` coincide (no measurability assumptions on `s` or `f`). The special case `t = univ` is especially useful. --- src/analysis/convolution.lean | 2 +- .../constructions/borel_space.lean | 5 + .../conditional_expectation/basic.lean | 2 +- src/measure_theory/function/l1_space.lean | 16 ++ .../function/locally_integrable.lean | 2 +- .../function/strongly_measurable/basic.lean | 37 ++- src/measure_theory/group/arithmetic.lean | 11 + .../integral/integrable_on.lean | 77 +++++- src/measure_theory/integral/set_integral.lean | 219 +++++++++++++++--- src/measure_theory/measure/measure_space.lean | 88 +++++-- 10 files changed, 394 insertions(+), 65 deletions(-) diff --git a/src/analysis/convolution.lean b/src/analysis/convolution.lean index be505012c1379..c46120ed87890 100644 --- a/src/analysis/convolution.lean +++ b/src/analysis/convolution.lean @@ -211,7 +211,7 @@ lemma bdd_above.convolution_exists_at' {x₀ : G} (hf : integrable_on f s μ) (hmg : ae_strongly_measurable g $ map (λ t, x₀ - t) (μ.restrict s)) : convolution_exists_at f g x₀ L μ := begin - rw [convolution_exists_at, ← integrable_on_iff_integrable_of_support_subset h2s hs], + rw [convolution_exists_at, ← integrable_on_iff_integrable_of_support_subset h2s], set s' := (λ t, - t + x₀) ⁻¹' s, have : ∀ᵐ (t : G) ∂(μ.restrict s), ‖L (f t) (g (x₀ - t))‖ ≤ s.indicator (λ t, ‖L‖ * ‖f t‖ * ⨆ i : s', ‖g i‖) t, diff --git a/src/measure_theory/constructions/borel_space.lean b/src/measure_theory/constructions/borel_space.lean index 73c77a38e0bf5..a1539b3273349 100644 --- a/src/measure_theory/constructions/borel_space.lean +++ b/src/measure_theory/constructions/borel_space.lean @@ -480,6 +480,11 @@ lemma measurable_set_lt [second_countable_topology α] {f g : δ → α} (hf : m (hg : measurable g) : measurable_set {a | f a < g a} := hf.prod_mk hg measurable_set_lt' +lemma null_measurable_set_lt [second_countable_topology α] {μ : measure δ} {f g : δ → α} + (hf : ae_measurable f μ) (hg : ae_measurable g μ) : + null_measurable_set {a | f a < g a} μ := +(hf.prod_mk hg).null_measurable measurable_set_lt' + lemma set.ord_connected.measurable_set (h : ord_connected s) : measurable_set s := begin let u := ⋃ (x ∈ s) (y ∈ s), Ioo x y, diff --git a/src/measure_theory/function/conditional_expectation/basic.lean b/src/measure_theory/function/conditional_expectation/basic.lean index 3e39ed9c4daae..4ad8e7062a3e9 100644 --- a/src/measure_theory/function/conditional_expectation/basic.lean +++ b/src/measure_theory/function/conditional_expectation/basic.lean @@ -856,7 +856,7 @@ lemma integral_norm_le_of_forall_fin_meas_integral_eq (hm : m ≤ m0) {f g : α (hs : measurable_set[m] s) (hμs : μ s ≠ ∞) : ∫ x in s, ‖g x‖ ∂μ ≤ ∫ x in s, ‖f x‖ ∂μ := begin - rw [integral_norm_eq_pos_sub_neg (hg.mono hm) hgi, integral_norm_eq_pos_sub_neg hf hfi], + rw [integral_norm_eq_pos_sub_neg hgi, integral_norm_eq_pos_sub_neg hfi], have h_meas_nonneg_g : measurable_set[m] {x | 0 ≤ g x}, from (@strongly_measurable_const _ _ m _ _).measurable_set_le hg, have h_meas_nonneg_f : measurable_set {x | 0 ≤ f x}, diff --git a/src/measure_theory/function/l1_space.lean b/src/measure_theory/function/l1_space.lean index 062832af9c13a..b63912f5a5cf6 100644 --- a/src/measure_theory/function/l1_space.lean +++ b/src/measure_theory/function/l1_space.lean @@ -672,6 +672,22 @@ lemma mem_ℒp.integrable {q : ℝ≥0∞} (hq1 : 1 ≤ q) {f : α → β} [is_f (hfq : mem_ℒp f q μ) : integrable f μ := mem_ℒp_one_iff_integrable.mp (hfq.mem_ℒp_of_exponent_le hq1) +/-- A non-quantitative version of Markov inequality for integrable functions: the measure of points +where `‖f x‖ ≥ ε` is finite for all positive `ε`. -/ +lemma integrable.measure_ge_lt_top {f : α → β} (hf : integrable f μ) {ε : ℝ} (hε : 0 < ε) : + μ {x | ε ≤ ‖f x‖} < ∞ := +begin + rw show {x | ε ≤ ‖f x‖} = {x | ennreal.of_real ε ≤ ‖f x‖₊}, + by simp only [ennreal.of_real, real.to_nnreal_le_iff_le_coe, ennreal.coe_le_coe, coe_nnnorm], + refine (meas_ge_le_mul_pow_snorm μ one_ne_zero ennreal.one_ne_top hf.1 _).trans_lt _, + { simpa only [ne.def, ennreal.of_real_eq_zero, not_le] using hε }, + apply ennreal.mul_lt_top, + { simpa only [ennreal.one_to_real, ennreal.rpow_one, ne.def, ennreal.inv_eq_top, + ennreal.of_real_eq_zero, not_le] using hε }, + simpa only [ennreal.one_to_real, ennreal.rpow_one] + using (mem_ℒp_one_iff_integrable.2 hf).snorm_ne_top, +end + lemma lipschitz_with.integrable_comp_iff_of_antilipschitz {K K'} {f : α → β} {g : β → γ} (hg : lipschitz_with K g) (hg' : antilipschitz_with K' g) (g0 : g 0 = 0) : integrable (g ∘ f) μ ↔ integrable f μ := diff --git a/src/measure_theory/function/locally_integrable.lean b/src/measure_theory/function/locally_integrable.lean index 520eb5619f3dd..54eff2ef7aa62 100644 --- a/src/measure_theory/function/locally_integrable.lean +++ b/src/measure_theory/function/locally_integrable.lean @@ -231,7 +231,7 @@ hf.integrable_on_Ioc /-- A continuous function with compact support is integrable on the whole space. -/ lemma continuous.integrable_of_has_compact_support (hf : continuous f) (hcf : has_compact_support f) : integrable f μ := -(integrable_on_iff_integrable_of_support_subset (subset_tsupport f) measurable_set_closure).mp $ +(integrable_on_iff_integrable_of_support_subset (subset_tsupport f)).mp $ hf.continuous_on.integrable_on_compact hcf end borel diff --git a/src/measure_theory/function/strongly_measurable/basic.lean b/src/measure_theory/function/strongly_measurable/basic.lean index d5414eb8b9c20..1e406c67fbf5f 100644 --- a/src/measure_theory/function/strongly_measurable/basic.lean +++ b/src/measure_theory/function/strongly_measurable/basic.lean @@ -113,7 +113,7 @@ open_locale measure_theory /-! ## Strongly measurable functions -/ -lemma strongly_measurable.ae_strongly_measurable {α β} {m0 : measurable_space α} +protected lemma strongly_measurable.ae_strongly_measurable {α β} {m0 : measurable_space α} [topological_space β] {f : α → β} {μ : measure α} (hf : strongly_measurable f) : ae_strongly_measurable f μ := ⟨f, hf, eventually_eq.refl _ _⟩ @@ -1425,6 +1425,41 @@ protected lemma indicator [has_zero β] ae_strongly_measurable (s.indicator f) μ := (ae_strongly_measurable_indicator_iff hs).mpr hfm.restrict +lemma null_measurable_set_eq_fun {E} [topological_space E] [metrizable_space E] + {f g : α → E} (hf : ae_strongly_measurable f μ) (hg : ae_strongly_measurable g μ) : + null_measurable_set {x | f x = g x} μ := +begin + apply (hf.strongly_measurable_mk.measurable_set_eq_fun hg.strongly_measurable_mk) + .null_measurable_set.congr, + filter_upwards [hf.ae_eq_mk, hg.ae_eq_mk] with x hfx hgx, + change (hf.mk f x = hg.mk g x) = (f x = g x), + simp only [hfx, hgx] +end + +lemma null_measurable_set_lt + [linear_order β] [order_closed_topology β] [pseudo_metrizable_space β] + {f g : α → β} (hf : ae_strongly_measurable f μ) + (hg : ae_strongly_measurable g μ) : + null_measurable_set {a | f a < g a} μ := +begin + apply (hf.strongly_measurable_mk.measurable_set_lt hg.strongly_measurable_mk) + .null_measurable_set.congr, + filter_upwards [hf.ae_eq_mk, hg.ae_eq_mk] with x hfx hgx, + change (hf.mk f x < hg.mk g x) = (f x < g x), + simp only [hfx, hgx] +end + +lemma null_measurable_set_le [preorder β] [order_closed_topology β] [pseudo_metrizable_space β] + {f g : α → β} (hf : ae_strongly_measurable f μ) (hg : ae_strongly_measurable g μ) : + null_measurable_set {a | f a ≤ g a} μ := +begin + apply (hf.strongly_measurable_mk.measurable_set_le hg.strongly_measurable_mk) + .null_measurable_set.congr, + filter_upwards [hf.ae_eq_mk, hg.ae_eq_mk] with x hfx hgx, + change (hf.mk f x ≤ hg.mk g x) = (f x ≤ g x), + simp only [hfx, hgx] +end + lemma _root_.ae_strongly_measurable_of_ae_strongly_measurable_trim {α} {m m0 : measurable_space α} {μ : measure α} (hm : m ≤ m0) {f : α → β} (hf : ae_strongly_measurable f (μ.trim hm)) : ae_strongly_measurable f μ := diff --git a/src/measure_theory/group/arithmetic.lean b/src/measure_theory/group/arithmetic.lean index a2ede97a8e05a..3f35f14bcb471 100644 --- a/src/measure_theory/group/arithmetic.lean +++ b/src/measure_theory/group/arithmetic.lean @@ -336,6 +336,17 @@ begin simp_rw [set.mem_set_of_eq, pi.sub_apply, sub_eq_zero], end +lemma null_measurable_set_eq_fun {E} [measurable_space E] [add_group E] + [measurable_singleton_class E] [has_measurable_sub₂ E] {f g : α → E} + (hf : ae_measurable f μ) (hg : ae_measurable g μ) : + null_measurable_set {x | f x = g x} μ := +begin + apply (measurable_set_eq_fun hf.measurable_mk hg.measurable_mk).null_measurable_set.congr, + filter_upwards [hf.ae_eq_mk, hg.ae_eq_mk] with x hfx hgx, + change (hf.mk f x = hg.mk g x) = (f x = g x), + simp only [hfx, hgx], +end + lemma measurable_set_eq_fun_of_countable {m : measurable_space α} {E} [measurable_space E] [measurable_singleton_class E] [countable E] {f g : α → E} (hf : measurable f) (hg : measurable g) : diff --git a/src/measure_theory/integral/integrable_on.lean b/src/measure_theory/integral/integrable_on.lean index 1275358d89507..a1edc3eddc9b1 100644 --- a/src/measure_theory/integral/integrable_on.lean +++ b/src/measure_theory/integral/integrable_on.lean @@ -231,12 +231,83 @@ begin simpa only [set.univ_inter, measurable_set.univ, measure.restrict_apply] using hμs, end -lemma integrable_on_iff_integrable_of_support_subset {f : α → E} {s : set α} - (h1s : support f ⊆ s) (h2s : measurable_set s) : +/-- If a function is integrable on a set `s` and nonzero there, then the measurable hull of `s` is +well behaved: the restriction of the measure to `to_measurable μ s` coincides with its restriction +to `s`. -/ +lemma integrable_on.restrict_to_measurable (hf : integrable_on f s μ) (h's : ∀ x ∈ s, f x ≠ 0) : + μ.restrict (to_measurable μ s) = μ.restrict s := +begin + rcases exists_seq_strict_anti_tendsto (0 : ℝ) with ⟨u, u_anti, u_pos, u_lim⟩, + let v := λ n, to_measurable (μ.restrict s) {x | u n ≤ ‖f x‖}, + have A : ∀ n, μ (s ∩ v n) ≠ ∞, + { assume n, + rw [inter_comm, ← measure.restrict_apply (measurable_set_to_measurable _ _), + measure_to_measurable], + exact (hf.measure_ge_lt_top (u_pos n)).ne }, + apply measure.restrict_to_measurable_of_cover _ A, + assume x hx, + have : 0 < ‖f x‖, by simp only [h's x hx, norm_pos_iff, ne.def, not_false_iff], + obtain ⟨n, hn⟩ : ∃ n, u n < ‖f x‖, from ((tendsto_order.1 u_lim).2 _ this).exists, + refine mem_Union.2 ⟨n, _⟩, + exact subset_to_measurable _ _ hn.le +end + +/-- If a function is integrable on a set `s`, and vanishes on `t \ s`, then it is integrable on `t` +if `t` is null-measurable. -/ +lemma integrable_on.of_ae_diff_eq_zero (hf : integrable_on f s μ) + (ht : null_measurable_set t μ) (h't : ∀ᵐ x ∂μ, x ∈ t \ s → f x = 0) : + integrable_on f t μ := +begin + let u := {x ∈ s | f x ≠ 0}, + have hu : integrable_on f u μ := hf.mono_set (λ x hx, hx.1), + let v := to_measurable μ u, + have A : integrable_on f v μ, + { rw [integrable_on, hu.restrict_to_measurable], + { exact hu }, + { assume x hx, exact hx.2 } }, + have B : integrable_on f (t \ v) μ, + { apply integrable_on_zero.congr, + filter_upwards [ae_restrict_of_ae h't, ae_restrict_mem₀ + (ht.diff (measurable_set_to_measurable μ u).null_measurable_set)] with x hxt hx, + by_cases h'x : x ∈ s, + { by_contra H, + exact hx.2 (subset_to_measurable μ u ⟨h'x, ne.symm H⟩) }, + { exact (hxt ⟨hx.1, h'x⟩).symm, } }, + apply (A.union B).mono_set _, + rw union_diff_self, + exact subset_union_right _ _ +end + +/-- If a function is integrable on a set `s`, and vanishes on `t \ s`, then it is integrable on `t` +if `t` is measurable. -/ +lemma integrable_on.of_forall_diff_eq_zero (hf : integrable_on f s μ) + (ht : measurable_set t) (h't : ∀ x ∈ t \ s, f x = 0) : + integrable_on f t μ := +hf.of_ae_diff_eq_zero ht.null_measurable_set (eventually_of_forall h't) + +/-- If a function is integrable on a set `s` and vanishes almost everywhere on its complement, +then it is integrable. -/ +lemma integrable_on.integrable_of_ae_not_mem_eq_zero (hf : integrable_on f s μ) + (h't : ∀ᵐ x ∂μ, x ∉ s → f x = 0) : integrable f μ := +begin + rw ← integrable_on_univ, + apply hf.of_ae_diff_eq_zero null_measurable_set_univ, + filter_upwards [h't] with x hx h'x using hx h'x.2, +end + +/-- If a function is integrable on a set `s` and vanishes everywhere on its complement, +then it is integrable. -/ +lemma integrable_on.integrable_of_forall_not_mem_eq_zero (hf : integrable_on f s μ) + (h't : ∀ x ∉ s, f x = 0) : integrable f μ := +hf.integrable_of_ae_not_mem_eq_zero (eventually_of_forall (λ x hx, h't x hx)) + +lemma integrable_on_iff_integrable_of_support_subset (h1s : support f ⊆ s) : integrable_on f s μ ↔ integrable f μ := begin refine ⟨λ h, _, λ h, h.integrable_on⟩, - rwa [← indicator_eq_self.2 h1s, integrable_indicator_iff h2s] + apply h.integrable_of_forall_not_mem_eq_zero (λ x hx, _), + contrapose! hx, + exact h1s (mem_support.2 hx), end lemma integrable_on_Lp_of_measure_ne_top {E} [normed_add_comm_group E] diff --git a/src/measure_theory/integral/set_integral.lean b/src/measure_theory/integral/set_integral.lean index 6d8721454d123..85c9c82763295 100644 --- a/src/measure_theory/integral/set_integral.lean +++ b/src/measure_theory/integral/set_integral.lean @@ -61,10 +61,18 @@ variables [normed_add_comm_group E] {f g : α → E} {s t : set α} {μ ν : me variables [complete_space E] [normed_space ℝ E] +lemma set_integral_congr_ae₀ (hs : null_measurable_set s μ) (h : ∀ᵐ x ∂μ, x ∈ s → f x = g x) : + ∫ x in s, f x ∂μ = ∫ x in s, g x ∂μ := +integral_congr_ae ((ae_restrict_iff'₀ hs).2 h) + lemma set_integral_congr_ae (hs : measurable_set s) (h : ∀ᵐ x ∂μ, x ∈ s → f x = g x) : ∫ x in s, f x ∂μ = ∫ x in s, g x ∂μ := integral_congr_ae ((ae_restrict_iff' hs).2 h) +lemma set_integral_congr₀ (hs : null_measurable_set s μ) (h : eq_on f g s) : + ∫ x in s, f x ∂μ = ∫ x in s, g x ∂μ := +set_integral_congr_ae₀ hs $ eventually_of_forall h + lemma set_integral_congr (hs : measurable_set s) (h : eq_on f g s) : ∫ x in s, f x ∂μ = ∫ x in s, g x ∂μ := set_integral_congr_ae hs $ eventually_of_forall h @@ -83,14 +91,25 @@ lemma integral_union (hst : disjoint s t) (ht : measurable_set t) ∫ x in s ∪ t, f x ∂μ = ∫ x in s, f x ∂μ + ∫ x in t, f x ∂μ := integral_union_ae hst.ae_disjoint ht.null_measurable_set hfs hft -lemma integral_diff (ht : measurable_set t) (hfs : integrable_on f s μ) - (hft : integrable_on f t μ) (hts : t ⊆ s) : +lemma integral_diff (ht : measurable_set t) (hfs : integrable_on f s μ) (hts : t ⊆ s) : ∫ x in s \ t, f x ∂μ = ∫ x in s, f x ∂μ - ∫ x in t, f x ∂μ := begin rw [eq_sub_iff_add_eq, ← integral_union, diff_union_of_subset hts], - exacts [disjoint_sdiff_self_left, ht, hfs.mono_set (diff_subset _ _), hft] + exacts [disjoint_sdiff_self_left, ht, hfs.mono_set (diff_subset _ _), hfs.mono_set hts] end +lemma integral_inter_add_diff₀ (ht : null_measurable_set t μ) (hfs : integrable_on f s μ) : + ∫ x in s ∩ t, f x ∂μ + ∫ x in s \ t, f x ∂μ = ∫ x in s, f x ∂μ := +begin + rw [← measure.restrict_inter_add_diff₀ s ht, integral_add_measure], + { exact integrable.mono_measure hfs (measure.restrict_mono (inter_subset_left _ _) le_rfl) }, + { exact integrable.mono_measure hfs (measure.restrict_mono (diff_subset _ _) le_rfl) } +end + +lemma integral_inter_add_diff (ht : measurable_set t) (hfs : integrable_on f s μ) : + ∫ x in s ∩ t, f x ∂μ + ∫ x in s \ t, f x ∂μ = ∫ x in s, f x ∂μ := +integral_inter_add_diff₀ ht.null_measurable_set hfs + lemma integral_finset_bUnion {ι : Type*} (t : finset ι) {s : ι → set α} (hs : ∀ i ∈ t, measurable_set (s i)) (h's : set.pairwise ↑t (disjoint on s)) (hf : ∀ i ∈ t, integrable_on f (s i) μ) : @@ -121,11 +140,15 @@ lemma integral_empty : ∫ x in ∅, f x ∂μ = 0 := by rw [measure.restrict_em lemma integral_univ : ∫ x in univ, f x ∂μ = ∫ x, f x ∂μ := by rw [measure.restrict_univ] -lemma integral_add_compl (hs : measurable_set s) (hfi : integrable f μ) : +lemma integral_add_compl₀ (hs : null_measurable_set s μ) (hfi : integrable f μ) : ∫ x in s, f x ∂μ + ∫ x in sᶜ, f x ∂μ = ∫ x, f x ∂μ := -by rw [← integral_union (@disjoint_compl_right (set α) _ _) hs.compl +by rw [← integral_union_ae (@disjoint_compl_right (set α) _ _).ae_disjoint hs.compl hfi.integrable_on hfi.integrable_on, union_compl_self, integral_univ] +lemma integral_add_compl (hs : measurable_set s) (hfi : integrable f μ) : + ∫ x in s, f x ∂μ + ∫ x in sᶜ, f x ∂μ = ∫ x, f x ∂μ := +integral_add_compl₀ hs.null_measurable_set hfi + /-- For a function `f` and a measurable set `s`, the integral of `indicator s f` over the whole space is equal to `∫ x in s, f x ∂μ` defined as `∫ x, f x ∂(μ.restrict s)`. -/ lemma integral_indicator (hs : measurable_set s) : @@ -165,14 +188,14 @@ lemma of_real_set_integral_one {α : Type*} {m : measurable_space α} (μ : meas of_real_set_integral_one_of_measure_ne_top (measure_ne_top μ s) lemma integral_piecewise [decidable_pred (∈ s)] (hs : measurable_set s) - {f g : α → E} (hf : integrable_on f s μ) (hg : integrable_on g sᶜ μ) : + (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.integrable_indicator hs) (hg.integrable_indicator hs.compl), integral_indicator hs, integral_indicator hs.compl] lemma tendsto_set_integral_of_monotone {ι : Type*} [countable ι] [semilattice_sup ι] - {s : ι → set α} {f : α → E} (hsm : ∀ i, measurable_set (s i)) + {s : ι → set α} (hsm : ∀ i, measurable_set (s i)) (h_mono : monotone s) (hfi : integrable_on f (⋃ n, s n) μ) : tendsto (λ i, ∫ a in s i, f a ∂μ) at_top (𝓝 (∫ a in (⋃ n, s n), f a ∂μ)) := begin @@ -187,7 +210,7 @@ begin have : ∀ᶠ i in at_top, ν (s i) ∈ Icc (ν S - ε) (ν S + ε), from tendsto_measure_Union h_mono (ennreal.Icc_mem_nhds hfi'.ne (ennreal.coe_pos.2 ε0).ne'), refine this.mono (λ i hi, _), - rw [mem_closed_ball_iff_norm', ← integral_diff (hsm i) hfi (hfi.mono_set hsub) hsub, + rw [mem_closed_ball_iff_norm', ← integral_diff (hsm i) hfi hsub, ← coe_nnnorm, nnreal.coe_le_coe, ← ennreal.coe_le_coe], refine (ennnorm_integral_le_lintegral_ennnorm _).trans _, rw [← with_density_apply _ (hSm.diff (hsm _)), ← hν, measure_diff hsub (hsm _)], @@ -195,7 +218,7 @@ begin (hi.2.trans_lt $ ennreal.add_lt_top.2 ⟨hfi', ennreal.coe_lt_top⟩).ne] end -lemma has_sum_integral_Union_ae {ι : Type*} [countable ι] {s : ι → set α} {f : α → E} +lemma has_sum_integral_Union_ae {ι : Type*} [countable ι] {s : ι → set α} (hm : ∀ i, null_measurable_set (s i) μ) (hd : pairwise (ae_disjoint μ on s)) (hfi : integrable_on f (⋃ i, s i) μ) : has_sum (λ n, ∫ a in s n, f a ∂ μ) (∫ a in ⋃ n, s n, f a ∂μ) := @@ -204,65 +227,191 @@ begin exact has_sum_integral_measure hfi end -lemma has_sum_integral_Union {ι : Type*} [countable ι] {s : ι → set α} {f : α → E} +lemma has_sum_integral_Union {ι : Type*} [countable ι] {s : ι → set α} (hm : ∀ i, measurable_set (s i)) (hd : pairwise (disjoint on s)) (hfi : integrable_on f (⋃ i, s i) μ) : has_sum (λ n, ∫ a in s n, f a ∂ μ) (∫ a in ⋃ n, s n, f a ∂μ) := has_sum_integral_Union_ae (λ i, (hm i).null_measurable_set) (hd.mono (λ i j h, h.ae_disjoint)) hfi -lemma integral_Union {ι : Type*} [countable ι] {s : ι → set α} {f : α → E} +lemma integral_Union {ι : Type*} [countable ι] {s : ι → set α} (hm : ∀ i, measurable_set (s i)) (hd : pairwise (disjoint on s)) (hfi : integrable_on f (⋃ i, s i) μ) : (∫ a in (⋃ n, s n), f a ∂μ) = ∑' n, ∫ a in s n, f a ∂ μ := (has_sum.tsum_eq (has_sum_integral_Union hm hd hfi)).symm -lemma integral_Union_ae {ι : Type*} [countable ι] {s : ι → set α} {f : α → E} +lemma integral_Union_ae {ι : Type*} [countable ι] {s : ι → set α} (hm : ∀ i, null_measurable_set (s i) μ) (hd : pairwise (ae_disjoint μ on s)) (hfi : integrable_on f (⋃ i, s i) μ) : (∫ a in (⋃ n, s n), f a ∂μ) = ∑' n, ∫ a in s n, f a ∂ μ := (has_sum.tsum_eq (has_sum_integral_Union_ae hm hd hfi)).symm -lemma set_integral_eq_zero_of_forall_eq_zero {f : α → E} (hf : strongly_measurable f) - (ht_eq : ∀ x ∈ t, f x = 0) : +lemma set_integral_eq_zero_of_ae_eq_zero (ht_eq : ∀ᵐ x ∂μ, x ∈ t → f x = 0) : + ∫ x in t, f x ∂μ = 0 := +begin + by_cases hf : ae_strongly_measurable f (μ.restrict t), swap, + { rw integral_undef, + contrapose! hf, + exact hf.1 }, + have : ∫ x in t, hf.mk f x ∂μ = 0, + { refine integral_eq_zero_of_ae _, + rw [eventually_eq, ae_restrict_iff + (hf.strongly_measurable_mk.measurable_set_eq_fun strongly_measurable_zero)], + filter_upwards [ae_imp_of_ae_restrict hf.ae_eq_mk, ht_eq] with x hx h'x h''x, + rw ← hx h''x, + exact h'x h''x }, + rw ← this, + exact integral_congr_ae hf.ae_eq_mk, +end + +lemma set_integral_eq_zero_of_forall_eq_zero (ht_eq : ∀ x ∈ t, f x = 0) : ∫ x in t, f x ∂μ = 0 := +set_integral_eq_zero_of_ae_eq_zero (eventually_of_forall ht_eq) + +lemma integral_union_eq_left_of_ae_aux (ht_eq : ∀ᵐ x ∂(μ.restrict t), f x = 0) + (haux : strongly_measurable f) (H : integrable_on f (s ∪ t) μ) : + ∫ x in (s ∪ t), f x ∂μ = ∫ x in s, f x ∂μ := begin - refine integral_eq_zero_of_ae _, - rw [eventually_eq, ae_restrict_iff (hf.measurable_set_eq_fun strongly_measurable_zero)], - refine eventually_of_forall (λ x hx, _), - rw pi.zero_apply, - exact ht_eq x hx, + let k := f ⁻¹' {0}, + have hk : measurable_set k, + { borelize E, exact haux.measurable (measurable_set_singleton _) }, + have h's : integrable_on f s μ := H.mono (subset_union_left _ _) le_rfl, + have A : ∀ (u : set α), ∫ x in u ∩ k, f x ∂μ = 0 := + λ u, set_integral_eq_zero_of_forall_eq_zero (λ x hx, hx.2), + rw [← integral_inter_add_diff hk h's, ← integral_inter_add_diff hk H, A, A, zero_add, zero_add, + union_diff_distrib, union_comm], + apply set_integral_congr_set_ae, + rw union_ae_eq_right, + apply measure_mono_null (diff_subset _ _), + rw measure_zero_iff_ae_nmem, + filter_upwards [ae_imp_of_ae_restrict ht_eq] with x hx h'x using h'x.2 (hx h'x.1), end -lemma set_integral_union_eq_left {f : α → E} (hf : strongly_measurable f) (hfi : integrable f μ) - (hs : measurable_set s) (ht_eq : ∀ x ∈ t, f x = 0) : +lemma integral_union_eq_left_of_ae (ht_eq : ∀ᵐ x ∂(μ.restrict t), f x = 0) : ∫ x in (s ∪ t), f x ∂μ = ∫ x in s, f x ∂μ := begin - rw [← set.union_diff_self, union_comm, integral_union, - set_integral_eq_zero_of_forall_eq_zero _ (λ x hx, ht_eq x (diff_subset _ _ hx)), zero_add], - exacts [hf, disjoint_sdiff_self_left, hs, hfi.integrable_on, hfi.integrable_on] + have ht : integrable_on f t μ, + { apply integrable_on.congr_fun' integrable_on_zero, symmetry, exact ht_eq }, + by_cases H : integrable_on f (s ∪ t) μ, swap, + { rw [integral_undef H, integral_undef], simpa [integrable_on_union, ht] using H }, + let f' := H.1.mk f, + calc ∫ (x : α) in s ∪ t, f x ∂μ + = ∫ (x : α) in s ∪ t, f' x ∂μ : integral_congr_ae H.1.ae_eq_mk + ... = ∫ x in s, f' x ∂μ : + begin + apply integral_union_eq_left_of_ae_aux _ H.1.strongly_measurable_mk + (H.congr_fun' H.1.ae_eq_mk), + filter_upwards [ht_eq, ae_mono (measure.restrict_mono (subset_union_right s t) le_rfl) + H.1.ae_eq_mk] with x hx h'x, + rw [← h'x, hx] + end + ... = ∫ x in s, f x ∂μ : integral_congr_ae + (ae_mono (measure.restrict_mono (subset_union_left s t) le_rfl) H.1.ae_eq_mk.symm) end -lemma set_integral_neg_eq_set_integral_nonpos [linear_order E] [order_closed_topology E] - {f : α → E} (hf : strongly_measurable f) (hfi : integrable f μ) : +lemma integral_union_eq_left_of_forall₀ {f : α → E} + (ht : null_measurable_set t μ) (ht_eq : ∀ x ∈ t, f x = 0) : + ∫ x in (s ∪ t), f x ∂μ = ∫ x in s, f x ∂μ := +integral_union_eq_left_of_ae ((ae_restrict_iff'₀ ht).2 (eventually_of_forall ht_eq)) + +lemma integral_union_eq_left_of_forall {f : α → E} + (ht : measurable_set t) (ht_eq : ∀ x ∈ t, f x = 0) : + ∫ x in (s ∪ t), f x ∂μ = ∫ x in s, f x ∂μ := +integral_union_eq_left_of_forall₀ ht.null_measurable_set ht_eq + +lemma set_integral_eq_of_subset_of_ae_diff_eq_zero_aux + (hts : s ⊆ t) (h't : ∀ᵐ x ∂μ, x ∈ t \ s → f x = 0) + (haux : strongly_measurable f) (h'aux : integrable_on f t μ) : + ∫ x in t, f x ∂μ = ∫ x in s, f x ∂μ := +begin + let k := f ⁻¹' {0}, + have hk : measurable_set k, + { borelize E, exact haux.measurable (measurable_set_singleton _) }, + calc + ∫ x in t, f x ∂μ + = ∫ x in t ∩ k, f x ∂μ + ∫ x in t \ k, f x ∂μ : + by rw integral_inter_add_diff hk h'aux + ... = ∫ x in t \ k, f x ∂μ : + by { rw [set_integral_eq_zero_of_forall_eq_zero (λ x hx, _), zero_add], exact hx.2 } + ... = ∫ x in s \ k, f x ∂μ : + begin + apply set_integral_congr_set_ae, + filter_upwards [h't] with x hx, + change (x ∈ t \ k) = (x ∈ s \ k), + simp only [mem_preimage, mem_singleton_iff, eq_iff_iff, and.congr_left_iff, mem_diff], + assume h'x, + by_cases xs : x ∈ s, + { simp only [xs, hts xs] }, + { simp only [xs, iff_false], + assume xt, + exact h'x (hx ⟨xt, xs⟩) } + end + ... = ∫ x in s ∩ k, f x ∂μ + ∫ x in s \ k, f x ∂μ : + begin + have : ∀ x ∈ s ∩ k, f x = 0 := λ x hx, hx.2, + rw [set_integral_eq_zero_of_forall_eq_zero this, zero_add], + end + ... = ∫ x in s, f x ∂μ : by rw integral_inter_add_diff hk (h'aux.mono hts le_rfl) +end + +/-- If a function vanishes almost everywhere on `t \ s` with `s ⊆ t`, then its integrals on `s` +and `t` coincide if `t` is null-measurable. -/ +lemma set_integral_eq_of_subset_of_ae_diff_eq_zero + (ht : null_measurable_set t μ) (hts : s ⊆ t) (h't : ∀ᵐ x ∂μ, x ∈ t \ s → f x = 0) : + ∫ x in t, f x ∂μ = ∫ x in s, f x ∂μ := +begin + by_cases h : integrable_on f t μ, swap, + { have : ¬(integrable_on f s μ) := λ H, h (H.of_ae_diff_eq_zero ht h't), + rw [integral_undef h, integral_undef this] }, + let f' := h.1.mk f, + calc ∫ x in t, f x ∂μ + = ∫ x in t, f' x ∂μ : integral_congr_ae h.1.ae_eq_mk + ... = ∫ x in s, f' x ∂μ : + begin + apply set_integral_eq_of_subset_of_ae_diff_eq_zero_aux hts _ h.1.strongly_measurable_mk + (h.congr h.1.ae_eq_mk), + filter_upwards [h't, ae_imp_of_ae_restrict h.1.ae_eq_mk] with x hx h'x h''x, + rw [← h'x h''x.1, hx h''x] + end + ... = ∫ x in s, f x ∂μ : + begin + apply integral_congr_ae, + apply ae_restrict_of_ae_restrict_of_subset hts, + exact h.1.ae_eq_mk.symm + end +end + +/-- If a function vanishes on `t \ s` with `s ⊆ t`, then its integrals on `s` +and `t` coincide if `t` is measurable. -/ +lemma set_integral_eq_of_subset_of_forall_diff_eq_zero + (ht : measurable_set t) (hts : s ⊆ t) (h't : ∀ x ∈ t \ s, f x = 0) : + ∫ x in t, f x ∂μ = ∫ x in s, f x ∂μ := +set_integral_eq_of_subset_of_ae_diff_eq_zero ht.null_measurable_set hts + (eventually_of_forall (λ x hx, h't x hx)) + +lemma set_integral_neg_eq_set_integral_nonpos [linear_order E] + {f : α → E} (hf : ae_strongly_measurable f μ) : ∫ x in {x | f x < 0}, f x ∂μ = ∫ x in {x | f x ≤ 0}, f x ∂μ := begin have h_union : {x | f x ≤ 0} = {x | f x < 0} ∪ {x | f x = 0}, by { ext, simp_rw [set.mem_union, set.mem_set_of_eq], exact le_iff_lt_or_eq, }, rw h_union, - exact (set_integral_union_eq_left hf hfi (hf.measurable_set_lt strongly_measurable_const) - (λ x hx, hx)).symm, + have B : null_measurable_set {x | f x = 0} μ, + from hf.null_measurable_set_eq_fun ae_strongly_measurable_zero, + symmetry, + refine integral_union_eq_left_of_ae _, + filter_upwards [ae_restrict_mem₀ B] with x hx using hx, end -lemma integral_norm_eq_pos_sub_neg {f : α → ℝ} (hf : strongly_measurable f) - (hfi : integrable f μ) : +lemma integral_norm_eq_pos_sub_neg {f : α → ℝ} (hfi : integrable f μ) : ∫ x, ‖f x‖ ∂μ = ∫ x in {x | 0 ≤ f x}, f x ∂μ - ∫ x in {x | f x ≤ 0}, f x ∂μ := -have h_meas : measurable_set {x | 0 ≤ f x}, from strongly_measurable_const.measurable_set_le hf, +have h_meas : null_measurable_set {x | 0 ≤ f x} μ, + from ae_strongly_measurable_const.null_measurable_set_le hfi.1, calc ∫ x, ‖f x‖ ∂μ = ∫ x in {x | 0 ≤ f x}, ‖f x‖ ∂μ + ∫ x in {x | 0 ≤ f x}ᶜ, ‖f x‖ ∂μ : - by rw ← integral_add_compl h_meas hfi.norm + by rw ← integral_add_compl₀ h_meas hfi.norm ... = ∫ x in {x | 0 ≤ f x}, f x ∂μ + ∫ x in {x | 0 ≤ f x}ᶜ, ‖f x‖ ∂μ : begin congr' 1, - refine set_integral_congr h_meas (λ x hx, _), + refine set_integral_congr₀ h_meas (λ x hx, _), dsimp only, rw [real.norm_eq_abs, abs_eq_self.mpr _], exact hx, @@ -271,14 +420,14 @@ end begin congr' 1, rw ← integral_neg, - refine set_integral_congr h_meas.compl (λ x hx, _), + refine set_integral_congr₀ h_meas.compl (λ x hx, _), dsimp only, rw [real.norm_eq_abs, abs_eq_neg_self.mpr _], rw [set.mem_compl_iff, set.nmem_set_of_iff] at hx, linarith, end ... = ∫ x in {x | 0 ≤ f x}, f x ∂μ - ∫ x in {x | f x ≤ 0}, f x ∂μ : -by { rw ← set_integral_neg_eq_set_integral_nonpos hf hfi, congr, ext1 x, simp, } +by { rw ← set_integral_neg_eq_set_integral_nonpos hfi.1, congr, ext1 x, simp, } lemma set_integral_const (c : E) : ∫ x in s, c ∂μ = (μ s).to_real • c := by rw [integral_const, measure.restrict_apply_univ] diff --git a/src/measure_theory/measure/measure_space.lean b/src/measure_theory/measure/measure_space.lean index dca25b506f92e..b4388e38b58b2 100644 --- a/src/measure_theory/measure/measure_space.lean +++ b/src/measure_theory/measure/measure_space.lean @@ -2435,6 +2435,13 @@ lemma ae_restrict_of_ae {s : set α} {p : α → Prop} (h : ∀ᵐ x ∂μ, p x) (∀ᵐ x ∂(μ.restrict s), p x) := eventually.filter_mono (ae_mono measure.restrict_le_self) h +lemma ae_restrict_iff'₀ {p : α → Prop} (hs : null_measurable_set s μ) : + (∀ᵐ x ∂(μ.restrict s), p x) ↔ ∀ᵐ x ∂μ, x ∈ s → p x := +begin + refine ⟨λ h, ae_imp_of_ae_restrict h, λ h, _⟩, + filter_upwards [ae_restrict_mem₀ hs, ae_restrict_of_ae h] with x hx h'x using h'x hx, +end + lemma ae_restrict_of_ae_restrict_of_subset {s t : set α} {p : α → Prop} (hst : s ⊆ t) (h : ∀ᵐ x ∂(μ.restrict t), p x) : (∀ᵐ x ∂(μ.restrict s), p x) := @@ -3149,12 +3156,12 @@ begin (λ b, g_mble (‹measurable_singleton_class β›.measurable_set_singleton b)) level_sets_disjoint, end -/-- The measurable superset `to_measurable μ t` of `t` (which has the same measure as `t`) -satisfies, for any measurable set `s`, the equality `μ (to_measurable μ t ∩ s) = μ (t ∩ s)`. -This only holds when `μ` is σ-finite. For a version without this assumption (but requiring -that `t` has finite measure), see `measure_to_measurable_inter`. -/ -lemma measure_to_measurable_inter_of_sigma_finite - [sigma_finite μ] {s : set α} (hs : measurable_set s) (t : set α) : +/-- If a set `t` is covered by a countable family of finite measure sets, then its measurable +superset `to_measurable μ t` (which has the same measure as `t`) satisfies, +for any measurable set `s`, the equality `μ (to_measurable μ t ∩ s) = μ (t ∩ s)`. -/ +lemma measure_to_measurable_inter_of_cover + {s : set α} (hs : measurable_set s) {t : set α} {v : ℕ → set α} (hv : t ⊆ ⋃ n, v n) + (h'v : ∀ n, μ (t ∩ v n) ≠ ∞) : μ (to_measurable μ t ∩ s) = μ (t ∩ s) := begin -- we show that there is a measurable superset of `t` satisfying the conclusion for any @@ -3162,43 +3169,57 @@ begin -- (which is well behaved for finite measure sets thanks to `measure_to_measurable_inter`), and -- the desired property passes to the union. have A : ∃ t' ⊇ t, measurable_set t' ∧ (∀ u, measurable_set u → μ (t' ∩ u) = μ (t ∩ u)), - { set t' := ⋃ n, to_measurable μ (t ∩ disjointed (spanning_sets μ) n) with ht', + { let w := λ n, to_measurable μ (t ∩ v n), + have hw : ∀ n, μ (w n) < ∞, + { assume n, + simp_rw [w, measure_to_measurable], + exact (h'v n).lt_top }, + set t' := ⋃ n, to_measurable μ (t ∩ disjointed w n) with ht', have tt' : t ⊆ t' := calc - t ⊆ ⋃ n, t ∩ disjointed (spanning_sets μ) n : - by rw [← inter_Union, Union_disjointed, Union_spanning_sets, inter_univ] - ... ⊆ ⋃ n, to_measurable μ (t ∩ disjointed (spanning_sets μ) n) : + t ⊆ ⋃ n, t ∩ disjointed w n : + begin + rw [← inter_Union, Union_disjointed, inter_Union], + assume x hx, + rcases mem_Union.1 (hv hx) with ⟨n, hn⟩, + refine mem_Union.2 ⟨n, _⟩, + have : x ∈ t ∩ v n := ⟨hx, hn⟩, + exact ⟨hx, subset_to_measurable μ _ this⟩, + end + ... ⊆ ⋃ n, to_measurable μ (t ∩ disjointed w n) : Union_mono (λ n, subset_to_measurable _ _), refine ⟨t', tt', measurable_set.Union (λ n, measurable_set_to_measurable μ _), λ u hu, _⟩, apply le_antisymm _ (measure_mono (inter_subset_inter tt' subset.rfl)), - calc μ (t' ∩ u) ≤ ∑' n, μ (to_measurable μ (t ∩ disjointed (spanning_sets μ) n) ∩ u) : + calc μ (t' ∩ u) ≤ ∑' n, μ (to_measurable μ (t ∩ disjointed w n) ∩ u) : by { rw [ht', Union_inter], exact measure_Union_le _ } - ... = ∑' n, μ ((t ∩ disjointed (spanning_sets μ) n) ∩ u) : + ... = ∑' n, μ ((t ∩ disjointed w n) ∩ u) : begin congr' 1, ext1 n, apply measure_to_measurable_inter hu, apply ne_of_lt, - calc μ (t ∩ disjointed (spanning_sets μ) n) - ≤ μ (disjointed (spanning_sets μ) n) : measure_mono (inter_subset_right _ _) - ... ≤ μ (spanning_sets μ n) : measure_mono (disjointed_le (spanning_sets μ) n) - ... < ∞ : measure_spanning_sets_lt_top _ _ + calc μ (t ∩ disjointed w n) + ≤ μ (t ∩ w n) : measure_mono ((inter_subset_inter_right _ (disjointed_le w n))) + ... ≤ μ (w n) : measure_mono (inter_subset_right _ _) + ... < ∞ : hw n end - ... = ∑' n, μ.restrict (t ∩ u) (disjointed (spanning_sets μ) n) : + ... = ∑' n, μ.restrict (t ∩ u) (disjointed w n) : begin congr' 1, ext1 n, rw [restrict_apply, inter_comm t _, inter_assoc], - exact measurable_set.disjointed (measurable_spanning_sets _) _ + apply measurable_set.disjointed (λ n, _), + exact measurable_set_to_measurable _ _ end - ... = μ.restrict (t ∩ u) (⋃ n, disjointed (spanning_sets μ) n) : + ... = μ.restrict (t ∩ u) (⋃ n, disjointed w n) : begin rw measure_Union, { exact disjoint_disjointed _ }, - { intro i, exact measurable_set.disjointed (measurable_spanning_sets _) _ } + { intro i, + apply measurable_set.disjointed (λ n, _), + exact measurable_set_to_measurable _ _ } end - ... = μ (t ∩ u) : - by rw [Union_disjointed, Union_spanning_sets, restrict_apply measurable_set.univ, - univ_inter] }, + ... ≤ μ.restrict (t ∩ u) univ : measure_mono (subset_univ _) + ... = μ (t ∩ u) : by rw [restrict_apply measurable_set.univ, univ_inter] }, -- thanks to the definition of `to_measurable`, the previous property will also be shared -- by `to_measurable μ t`, which is enough to conclude the proof. rw [to_measurable], @@ -3208,6 +3229,27 @@ begin { exact A.some_spec.snd.2 s hs }, end +lemma restrict_to_measurable_of_cover {s : set α} {v : ℕ → set α} (hv : s ⊆ ⋃ n, v n) + (h'v : ∀ n, μ (s ∩ v n) ≠ ∞) : + μ.restrict (to_measurable μ s) = μ.restrict s := +ext $ λ t ht, by simp only [restrict_apply ht, inter_comm t, + measure_to_measurable_inter_of_cover ht hv h'v] + +/-- The measurable superset `to_measurable μ t` of `t` (which has the same measure as `t`) +satisfies, for any measurable set `s`, the equality `μ (to_measurable μ t ∩ s) = μ (t ∩ s)`. +This only holds when `μ` is σ-finite. For a version without this assumption (but requiring +that `t` has finite measure), see `measure_to_measurable_inter`. -/ +lemma measure_to_measurable_inter_of_sigma_finite + [sigma_finite μ] {s : set α} (hs : measurable_set s) (t : set α) : + μ (to_measurable μ t ∩ s) = μ (t ∩ s) := +begin + have : t ⊆ ⋃ n, spanning_sets μ n, + { rw Union_spanning_sets, exact subset_univ _ }, + apply measure_to_measurable_inter_of_cover hs this (λ n, ne_of_lt _), + calc μ (t ∩ spanning_sets μ n) ≤ μ(spanning_sets μ n) : measure_mono (inter_subset_right _ _) + ... < ∞ : measure_spanning_sets_lt_top μ n, +end + @[simp] lemma restrict_to_measurable_of_sigma_finite [sigma_finite μ] (s : set α) : μ.restrict (to_measurable μ s) = μ.restrict s := ext $ λ t ht, by simp only [restrict_apply ht, inter_comm t, From 24f7dbdf87d8ed21f3819ad931db5686e6d3519f Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 9 Feb 2023 15:58:16 +0000 Subject: [PATCH 0452/1291] feat(analysis/quaternion): add `complete_space`, `module.free`, and `module.finite` instances (#18347) The main trick here is showing that the quaternions are in isometric equivalence with 4D euclidean space. --- src/algebra/quaternion.lean | 67 ++++++++++++++++++++++++++++++++++-- src/analysis/quaternion.lean | 49 +++++++++++++++++++++++--- 2 files changed, 109 insertions(+), 7 deletions(-) diff --git a/src/algebra/quaternion.lean b/src/algebra/quaternion.lean index b9b6694cb7567..78174e14b61b6 100644 --- a/src/algebra/quaternion.lean +++ b/src/algebra/quaternion.lean @@ -4,6 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import algebra.algebra.equiv +import linear_algebra.finrank +import linear_algebra.free_module.basic +import linear_algebra.free_module.finite.basic import set_theory.cardinal.ordinal import tactic.ring_exp @@ -65,6 +68,17 @@ def equiv_prod {R : Type*} (c₁ c₂ : R) : ℍ[R, c₁, c₂] ≃ R × R × R left_inv := λ ⟨a₁, a₂, a₃, a₄⟩, rfl, right_inv := λ ⟨a₁, a₂, a₃, a₄⟩, rfl } +/-- The equivalence between a quaternion algebra over `R` and `fin 4 → R`. -/ +@[simps symm_apply] +def equiv_tuple {R : Type*} (c₁ c₂ : R) : ℍ[R, c₁, c₂] ≃ (fin 4 → R) := +{ to_fun := λ a, ![a.1, a.2, a.3, a.4], + inv_fun := λ a, ⟨a 0, a 1, a 2, a 3⟩, + left_inv := λ ⟨a₁, a₂, a₃, a₄⟩, rfl, + right_inv := λ f, by ext ⟨_, _|_|_|_|_|⟨⟩⟩; refl } + +@[simp] lemma equiv_tuple_apply {R : Type*} (c₁ c₂ : R) (x : ℍ[R, c₁, c₂]) : + equiv_tuple c₁ c₂ x = ![x.re, x.im_i, x.im_j, x.im_k] := rfl + @[simp] lemma mk.eta {R : Type*} {c₁ c₂} : ∀ a : ℍ[R, c₁, c₂], mk a.1 a.2 a.3 a.4 = a | ⟨a₁, a₂, a₃, a₄⟩ := rfl @@ -191,7 +205,7 @@ instance : algebra R ℍ[R, c₁, c₂] := lemma algebra_map_eq (r : R) : algebra_map R ℍ[R,c₁,c₂] r = ⟨r, 0, 0, 0⟩ := rfl section -variables (R c₁ c₂) +variables (c₁ c₂) /-- `quaternion_algebra.re` as a `linear_map`-/ @[simps] def re_lm : ℍ[R, c₁, c₂] →ₗ[R] R := @@ -209,6 +223,37 @@ variables (R c₁ c₂) @[simps] def im_k_lm : ℍ[R, c₁, c₂] →ₗ[R] R := { to_fun := im_k, map_add' := λ x y, rfl, map_smul' := λ r x, rfl } +/-- `quaternion_algebra.equiv_tuple` as a linear equivalence. -/ +def linear_equiv_tuple : ℍ[R,c₁,c₂] ≃ₗ[R] (fin 4 → R) := +linear_equiv.symm -- proofs are not `rfl` in the forward direction + { to_fun := (equiv_tuple c₁ c₂).symm, + inv_fun := equiv_tuple c₁ c₂, + map_add' := λ v₁ v₂, rfl, + map_smul' := λ v₁ v₂, rfl, + .. (equiv_tuple c₁ c₂).symm } + +@[simp] lemma coe_linear_equiv_tuple : ⇑(linear_equiv_tuple c₁ c₂) = equiv_tuple c₁ c₂ := rfl +@[simp] lemma coe_linear_equiv_tuple_symm : + ⇑(linear_equiv_tuple c₁ c₂).symm = (equiv_tuple c₁ c₂).symm := rfl + +/-- `ℍ[R, c₁, c₂]` has a basis over `R` given by `1`, `i`, `j`, and `k`. -/ +noncomputable def basis_one_i_j_k : basis (fin 4) R ℍ[R, c₁, c₂] := +basis.of_equiv_fun $ linear_equiv_tuple c₁ c₂ + +@[simp] lemma coe_basis_one_i_j_k_repr (q : ℍ[R, c₁, c₂]) : + ⇑((basis_one_i_j_k c₁ c₂).repr q) = ![q.re, q.im_i, q.im_j, q.im_k] := rfl + +instance : module.finite R ℍ[R, c₁, c₂] := module.finite.of_basis (basis_one_i_j_k c₁ c₂) +instance : module.free R ℍ[R, c₁, c₂] := module.free.of_basis (basis_one_i_j_k c₁ c₂) + +lemma dim_eq_four [strong_rank_condition R] : module.rank R ℍ[R, c₁, c₂] = 4 := +by { rw [dim_eq_card_basis (basis_one_i_j_k c₁ c₂), fintype.card_fin], norm_num } + +lemma finrank_eq_four [strong_rank_condition R] : finite_dimensional.finrank R ℍ[R, c₁, c₂] = 4 := +have cardinal.to_nat 4 = 4, + by rw [←cardinal.to_nat_cast 4, nat.cast_bit0, nat.cast_bit0, nat.cast_one], +by rw [finite_dimensional.finrank, dim_eq_four, this] + end @[norm_cast, simp] lemma coe_sub : ((x - y : R) : ℍ[R, c₁, c₂]) = x - y := @@ -344,10 +389,19 @@ def quaternion (R : Type*) [has_one R] [has_neg R] := quaternion_algebra R (-1) localized "notation (name := quaternion) `ℍ[` R `]` := quaternion R" in quaternion -/-- The equivalence between the quaternions over R and R × R × R × R. -/ +/-- The equivalence between the quaternions over `R` and `R × R × R × R`. -/ +@[simps] def quaternion.equiv_prod (R : Type*) [has_one R] [has_neg R] : ℍ[R] ≃ R × R × R × R := quaternion_algebra.equiv_prod _ _ +/-- The equivalence between the quaternions over `R` and `fin 4 → R`. -/ +@[simps symm_apply] +def quaternion.equiv_tuple (R : Type*) [has_one R] [has_neg R] : ℍ[R] ≃ (fin 4 → R) := +quaternion_algebra.equiv_tuple _ _ + +@[simp] lemma quaternion.equiv_tuple_apply (R : Type*) [has_one R] [has_neg R] (x : ℍ[R]) : + quaternion.equiv_tuple R x = ![x.re, x.im_i, x.im_j, x.im_k] := rfl + namespace quaternion variables {R : Type*} [comm_ring R] (r x y z : R) (a b c : ℍ[R]) @@ -445,6 +499,15 @@ lemma mul_coe_eq_smul : a * r = r • a := quaternion_algebra.mul_coe_eq_smul r lemma smul_coe : x • (y : ℍ[R]) = ↑(x * y) := quaternion_algebra.smul_coe x y +instance : module.finite R ℍ[R] := quaternion_algebra.module.finite _ _ +instance : module.free R ℍ[R] := quaternion_algebra.module.free _ _ + +lemma dim_eq_four [strong_rank_condition R] : module.rank R ℍ[R] = 4 := +quaternion_algebra.dim_eq_four _ _ + +lemma finrank_eq_four [strong_rank_condition R] : finite_dimensional.finrank R ℍ[R] = 4 := +quaternion_algebra.finrank_eq_four _ _ + /-- Quaternion conjugate. -/ def conj : ℍ[R] ≃ₗ[R] ℍ[R] := quaternion_algebra.conj diff --git a/src/analysis/quaternion.lean b/src/analysis/quaternion.lean index 3e3000dc5c7fa..4fd74563713f3 100644 --- a/src/analysis/quaternion.lean +++ b/src/analysis/quaternion.lean @@ -1,10 +1,11 @@ /- Copyright (c) 2020 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury Kudryashov +Authors: Yury Kudryashov, Eric Wieser -/ import algebra.quaternion import analysis.inner_product_space.basic +import analysis.inner_product_space.pi_L2 /-! # Quaternions as a normed algebra @@ -15,6 +16,8 @@ In this file we define the following structures on the space `ℍ := ℍ[ℝ]` o * normed ring; * normed space over `ℝ`. +We show that the norm on `ℍ[ℝ]` agrees with the euclidean norm of its components. + ## Notation The following notation is available with `open_locale quaternion`: @@ -29,8 +32,6 @@ quaternion, normed ring, normed space, normed algebra localized "notation (name := quaternion.real) `ℍ` := quaternion ℝ" in quaternion open_locale real_inner_product_space -noncomputable theory - namespace quaternion instance : has_inner ℝ ℍ := ⟨λ a b, (a * b.conj).re⟩ @@ -39,7 +40,7 @@ lemma inner_self (a : ℍ) : ⟪a, a⟫ = norm_sq a := rfl lemma inner_def (a b : ℍ) : ⟪a, b⟫ = (a * b.conj).re := rfl -instance : inner_product_space ℝ ℍ := +noncomputable instance : inner_product_space ℝ ℍ := inner_product_space.of_core { inner := has_inner.inner, conj_sym := λ x y, by simp [inner_def, mul_comm], @@ -65,7 +66,7 @@ noncomputable instance : normed_division_ring ℍ := norm_mul' := λ a b, by { simp only [norm_eq_sqrt_real_inner, inner_self, norm_sq.map_mul], exact real.sqrt_mul norm_sq_nonneg _ } } -noncomputable instance : normed_algebra ℝ ℍ := +instance : normed_algebra ℝ ℍ := { norm_smul_le := λ a x, (norm_smul a x).le, to_algebra := quaternion.algebra } @@ -95,4 +96,42 @@ def of_complex : ℂ →ₐ[ℝ] ℍ := @[simp] lemma coe_of_complex : ⇑of_complex = coe := rfl +/-- The norm of the components as a euclidean vector equals the norm of the quaternion. -/ +lemma norm_pi_Lp_equiv_symm_equiv_tuple (x : ℍ) : + ‖(pi_Lp.equiv 2 (λ _ : fin 4, _)).symm (equiv_tuple ℝ x)‖ = ‖x‖ := +begin + rw [norm_eq_sqrt_real_inner, norm_eq_sqrt_real_inner, inner_self, norm_sq_def', pi_Lp.inner_apply, + fin.sum_univ_four], + simp_rw [is_R_or_C.inner_apply, star_ring_end_apply, star_trivial, ←sq], + refl, +end + +/-- `quaternion_algebra.linear_equiv_tuple` as a `linear_isometry_equiv`. -/ +@[simps apply symm_apply] +noncomputable def linear_isometry_equiv_tuple : ℍ ≃ₗᵢ[ℝ] euclidean_space ℝ (fin 4) := +{ to_fun := λ a, (pi_Lp.equiv _ (λ _ : fin 4, _)).symm ![a.1, a.2, a.3, a.4], + inv_fun := λ a, ⟨a 0, a 1, a 2, a 3⟩, + norm_map' := norm_pi_Lp_equiv_symm_equiv_tuple, + ..(quaternion_algebra.linear_equiv_tuple (-1 : ℝ) (-1 : ℝ)).trans + (pi_Lp.linear_equiv 2 ℝ (λ _ : fin 4, ℝ)).symm } + +@[continuity] lemma continuous_re : continuous (λ q : ℍ, q.re) := +(continuous_apply 0).comp linear_isometry_equiv_tuple.continuous + +@[continuity] lemma continuous_im_i : continuous (λ q : ℍ, q.im_i) := +(continuous_apply 1).comp linear_isometry_equiv_tuple.continuous + +@[continuity] lemma continuous_im_j : continuous (λ q : ℍ, q.im_j) := +(continuous_apply 2).comp linear_isometry_equiv_tuple.continuous + +@[continuity] lemma continuous_im_k : continuous (λ q : ℍ, q.im_k) := +(continuous_apply 3).comp linear_isometry_equiv_tuple.continuous + +instance : complete_space ℍ := +begin + have : uniform_embedding linear_isometry_equiv_tuple.to_linear_equiv.to_equiv.symm := + linear_isometry_equiv_tuple.to_continuous_linear_equiv.symm.uniform_embedding, + exact (complete_space_congr this).1 (by apply_instance) +end + end quaternion From dde670c9a3f503647fd5bfdf1037bad526d3397a Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 9 Feb 2023 15:58:17 +0000 Subject: [PATCH 0453/1291] chore(topology/constructions): pi.single is continuous (#18412) Forward-ported in https://github.com/leanprover-community/mathlib4/pull/2176. I found this convenient for proving continuity of `![r, 0, 0, 0]` via `convert`. --- src/topology/constructions.lean | 8 +++++++- 1 file changed, 7 insertions(+), 1 deletion(-) diff --git a/src/topology/constructions.lean b/src/topology/constructions.lean index b59b18db550df..9de0f22983304 100644 --- a/src/topology/constructions.lean +++ b/src/topology/constructions.lean @@ -1028,11 +1028,17 @@ lemma continuous.update [decidable_eq ι] (hf : continuous f) (i : ι) {g : α continuous (λ a, update (f a) i (g a)) := continuous_iff_continuous_at.2 $ λ x, hf.continuous_at.update i hg.continuous_at -/-- `update f i x` is continuous in `(f, x)`. -/ +/-- `function.update f i x` is continuous in `(f, x)`. -/ @[continuity] lemma continuous_update [decidable_eq ι] (i : ι) : continuous (λ f : (Π j, π j) × π i, update f.1 i f.2) := continuous_fst.update i continuous_snd +/-- `pi.mul_single i x` is continuous in `x`. -/ +@[continuity, to_additive "`pi.single i x` is continuous in `x`."] +lemma continuous_mul_single [Π i, has_one (π i)] [decidable_eq ι] (i : ι) : + continuous (λ x, (pi.mul_single i x : Π i, π i)) := +continuous_const.update _ continuous_id + lemma filter.tendsto.fin_insert_nth {n} {π : fin (n + 1) → Type*} [Π i, topological_space (π i)] (i : fin (n + 1)) {f : β → π i} {l : filter β} {x : π i} (hf : tendsto f l (𝓝 x)) {g : β → Π j : fin n, π (i.succ_above j)} {y : Π j, π (i.succ_above j)} (hg : tendsto g l (𝓝 y)) : From d39590fc8728fbf6743249802486f8c91ffe07bc Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Thu, 9 Feb 2023 19:18:04 +0000 Subject: [PATCH 0454/1291] refactor(topology/sets/opens): use a `structure` for `opens` and `open_nhds_of` (#18409) Also review API. --- src/algebraic_geometry/AffineScheme.lean | 39 ++-- .../Gamma_Spec_adjunction.lean | 2 +- src/algebraic_geometry/Spec.lean | 6 +- src/algebraic_geometry/function_field.lean | 10 +- .../locally_ringed_space.lean | 9 +- .../locally_ringed_space/has_colimits.lean | 6 +- .../morphisms/quasi_compact.lean | 18 +- .../morphisms/quasi_separated.lean | 2 +- .../morphisms/ring_hom_properties.lean | 7 +- src/algebraic_geometry/open_immersion.lean | 36 ++-- src/algebraic_geometry/presheafed_space.lean | 4 +- .../presheafed_space/gluing.lean | 9 +- .../presheafed_space/has_colimits.lean | 6 +- .../prime_spectrum/basic.lean | 12 +- .../projective_spectrum/scheme.lean | 4 +- .../projective_spectrum/topology.lean | 6 +- src/algebraic_geometry/properties.lean | 21 +- src/algebraic_geometry/ringed_space.lean | 6 +- src/algebraic_geometry/stalks.lean | 2 +- src/algebraic_geometry/structure_sheaf.lean | 14 +- src/measure_theory/measure/content.lean | 20 +- src/measure_theory/measure/haar.lean | 18 +- src/tactic/auto_cases.lean | 9 +- src/topology/category/Top/opens.lean | 13 +- src/topology/continuous_function/ideals.lean | 2 +- .../continuous_function/t0_sierpinski.lean | 2 +- src/topology/gluing.lean | 2 +- src/topology/sets/opens.lean | 200 ++++++++++++------ src/topology/sheaves/presheaf.lean | 9 +- src/topology/sheaves/punit.lean | 2 +- .../sheaves/sheaf_condition/sites.lean | 8 +- src/topology/sheaves/stalks.lean | 10 +- 32 files changed, 297 insertions(+), 217 deletions(-) diff --git a/src/algebraic_geometry/AffineScheme.lean b/src/algebraic_geometry/AffineScheme.lean index a03d6f223f6a6..916bf40550d9d 100644 --- a/src/algebraic_geometry/AffineScheme.lean +++ b/src/algebraic_geometry/AffineScheme.lean @@ -172,7 +172,7 @@ lemma is_basis_affine_open (X : Scheme) : begin rw opens.is_basis_iff_nbhd, rintros U x (hU : x ∈ (U : set X.carrier)), - obtain ⟨S, hS, hxS, hSU⟩ := X.affine_basis_cover_is_basis.exists_subset_of_mem_open hU U.prop, + obtain ⟨S, hS, hxS, hSU⟩ := X.affine_basis_cover_is_basis.exists_subset_of_mem_open hU U.is_open, refine ⟨⟨S, X.affine_basis_cover_is_basis.is_open hS⟩, _, hxS, hSU⟩, rcases hS with ⟨i, rfl⟩, exact range_is_affine_open_of_open_immersion _, @@ -226,8 +226,7 @@ begin haveI : is_affine _ := hU, convert range_is_affine_open_of_open_immersion (X.of_restrict U.open_embedding ≫ f), ext1, - change f.1.base '' U.1 = set.range (f.1.base ∘ coe), - rw [set.range_comp, subtype.range_coe], + exact set.image_eq_range _ _ end lemma is_affine_open_iff_of_is_open_immersion {X Y : Scheme} (f : X ⟶ Y) [H : is_open_immersion f] @@ -258,7 +257,7 @@ end lemma Scheme.Spec_map_presheaf_map_eq_to_hom {X : Scheme} {U V : opens X.carrier} (h : U = V) (W) : (Scheme.Spec.map (X.presheaf.map (eq_to_hom h).op).op).val.c.app W = - eq_to_hom (by { cases h, dsimp, induction W using opposite.rec, congr, ext1, simpa }) := + eq_to_hom (by { cases h, induction W using opposite.rec, dsimp, simp, refl }) := begin have : Scheme.Spec.map (X.presheaf.map (𝟙 (op U))).op = 𝟙 _, { rw [X.presheaf.map_id, op_id, Scheme.Spec.map_id] }, @@ -345,7 +344,7 @@ begin (X.of_restrict (X.basic_open r).open_embedding) _).mp, delta PresheafedSpace.is_open_immersion.open_functor, dsimp, - rw [opens.functor_obj_map_obj, opens.open_embedding_obj_top, inf_comm, + erw [opens.functor_obj_map_obj, opens.open_embedding_obj_top, inf_comm, ← Scheme.basic_open_res _ _ (hom_of_le le_top).op], exact hU.basic_open_is_affine _, end @@ -379,10 +378,10 @@ begin split, { rintro ⟨_, ⟨x, rfl⟩, rfl⟩, refine ⟨_, ⟨_, ⟨x, rfl⟩, rfl⟩, _⟩, - exact congr_arg subtype.val (X.map_prime_spectrum_basic_open_of_affine x) }, + exact congr_arg opens.carrier (X.map_prime_spectrum_basic_open_of_affine x) }, { rintro ⟨_, ⟨_, ⟨x, rfl⟩, rfl⟩, rfl⟩, refine ⟨_, ⟨x, rfl⟩, _⟩, - exact congr_arg subtype.val (X.map_prime_spectrum_basic_open_of_affine x).symm } + exact congr_arg opens.carrier (X.map_prime_spectrum_basic_open_of_affine x).symm } end lemma is_affine_open.exists_basic_open_le {X : Scheme} {U : opens X.carrier} @@ -391,7 +390,7 @@ lemma is_affine_open.exists_basic_open_le {X : Scheme} {U : opens X.carrier} begin haveI : is_affine _ := hU, obtain ⟨_, ⟨_, ⟨r, rfl⟩, rfl⟩, h₁, h₂⟩ := (is_basis_basic_open (X.restrict U.open_embedding)) - .exists_subset_of_mem_open _ ((opens.map U.inclusion).obj V).prop, + .exists_subset_of_mem_open _ ((opens.map U.inclusion).obj V).is_open, swap, exact ⟨x, h⟩, have : U.open_embedding.is_open_map.functor.obj ((X.restrict U.open_embedding).basic_open r) = X.basic_open (X.presheaf.map (eq_to_hom U.open_embedding_obj_top.symm).op r), @@ -445,7 +444,7 @@ begin exact RingedSpace.basic_open_le _ _ }, apply_instance end -. + lemma is_localization_basic_open {X : Scheme} {U : opens X.carrier} (hU : is_affine_open U) (f : X.presheaf.obj (op U)) : is_localization.away f (X.presheaf.obj (op $ X.basic_open f)) := @@ -500,7 +499,7 @@ begin end lemma exists_basic_open_le_affine_inter {X : Scheme} {U V : opens X.carrier} - (hU : is_affine_open U) (hV : is_affine_open V) (x : X.carrier) (hx : x ∈ U ∩ V) : + (hU : is_affine_open U) (hV : is_affine_open V) (x : X.carrier) (hx : x ∈ U ⊓ V) : ∃ (f : X.presheaf.obj $ op U) (g : X.presheaf.obj $ op V), X.basic_open f = X.basic_open g ∧ x ∈ X.basic_open f := begin @@ -641,23 +640,22 @@ begin apply_fun set.image hU.from_Spec.1.base at h, rw [set.image_preimage_eq_inter_range, set.image_preimage_eq_inter_range, hU.from_Spec_range] at h, - simp only [set.inter_self, subtype.val_eq_coe, set.inter_eq_right_iff_subset] + simp only [set.inter_self, opens.carrier_eq_coe, set.inter_eq_right_iff_subset] at h, ext1, - refine le_antisymm _ h, - simp only [set.Union_subset_iff, set_coe.forall, opens.supr_def, set.le_eq_subset, - subtype.coe_mk], + refine set.subset.antisymm _ h, + simp only [set.Union_subset_iff, set_coe.forall, opens.coe_supr], intros x hx, exact X.basic_open_le x }, { simp only [opens.supr_def, subtype.coe_mk, set.preimage_Union, subtype.val_eq_coe], congr' 3, { ext1 x, - exact congr_arg subtype.val (hU.from_Spec_map_basic_open _) }, - { exact congr_arg subtype.val hU.from_Spec_base_preimage } }, - { simp only [subtype.val_eq_coe, prime_spectrum.basic_open_eq_zero_locus_compl], + exact congr_arg opens.carrier (hU.from_Spec_map_basic_open _) }, + { exact congr_arg opens.carrier hU.from_Spec_base_preimage } }, + { simp only [opens.carrier_eq_coe, prime_spectrum.basic_open_eq_zero_locus_compl], rw [← set.compl_Inter, set.compl_univ_iff, ← prime_spectrum.zero_locus_Union, ← prime_spectrum.zero_locus_empty_iff_eq_top, prime_spectrum.zero_locus_span], - simp only [set.Union_singleton_eq_range, subtype.range_coe_subtype, set.set_of_mem_eq] } + simp only [set.Union_singleton_eq_range, subtype.range_val_subtype, set.set_of_mem_eq] } end lemma is_affine_open.self_le_basic_open_union_iff {X : Scheme} {U : opens X.carrier} @@ -712,9 +710,8 @@ begin exact hf₂ x }, rw ← V.prop.self_le_basic_open_union_iff, intros x hx, - simp only [exists_prop, set.mem_Union, set.mem_range, set_coe.exists, opens.supr_def, - exists_exists_eq_and, opens.mem_coe, subtype.coe_mk], - refine ⟨_, hf₁ ⟨x, hx⟩⟩, + rw [supr_range', opens.mem_supr], + exact ⟨_, hf₁ ⟨x, hx⟩⟩ end end algebraic_geometry diff --git a/src/algebraic_geometry/Gamma_Spec_adjunction.lean b/src/algebraic_geometry/Gamma_Spec_adjunction.lean index ebf8d621b59e4..ebe6c4c29ea3f 100644 --- a/src/algebraic_geometry/Gamma_Spec_adjunction.lean +++ b/src/algebraic_geometry/Gamma_Spec_adjunction.lean @@ -90,7 +90,7 @@ abbreviation to_Γ_Spec_map_basic_open : opens X := /-- The preimage is the basic open in `X` defined by the same element `r`. -/ lemma to_Γ_Spec_map_basic_open_eq : X.to_Γ_Spec_map_basic_open r = X.to_RingedSpace.basic_open r := -subtype.eq (X.to_Γ_Spec_preim_basic_open_eq r) +opens.ext (X.to_Γ_Spec_preim_basic_open_eq r) /-- The map from the global sections `Γ(X)` to the sections on the (preimage of) a basic open. -/ abbreviation to_to_Γ_Spec_map_basic_open : diff --git a/src/algebraic_geometry/Spec.lean b/src/algebraic_geometry/Spec.lean index 95d525a67666d..2b6fe31315a2c 100644 --- a/src/algebraic_geometry/Spec.lean +++ b/src/algebraic_geometry/Spec.lean @@ -302,13 +302,13 @@ def to_pushforward_stalk_alg_hom : S →ₐ[R] (Spec.Top_map (algebra_map R S) _* (structure_sheaf S).1).stalk p := { commutes' := λ _, rfl, ..(structure_sheaf.to_pushforward_stalk (algebra_map R S) p) } -. lemma is_localized_module_to_pushforward_stalk_alg_hom_aux (y) : ∃ (x : S × p.as_ideal.prime_compl), x.2 • y = to_pushforward_stalk_alg_hom R S p x.1 := begin obtain ⟨U, hp, s, e⟩ := Top.presheaf.germ_exist _ _ y, - obtain ⟨_, ⟨r, rfl⟩, hpr, hrU⟩ := prime_spectrum.is_topological_basis_basic_opens - .exists_subset_of_mem_open (show p ∈ U.1, from hp) U.2, + obtain ⟨_, ⟨r, rfl⟩, hpr : p ∈ prime_spectrum.basic_open r, + hrU : prime_spectrum.basic_open r ≤ U⟩ := prime_spectrum.is_topological_basis_basic_opens + .exists_subset_of_mem_open (show p ∈ ↑U, from hp) U.2, change prime_spectrum.basic_open r ≤ U at hrU, replace e := ((Spec.Top_map (algebra_map R S) _* (structure_sheaf S).1) .germ_res_apply (hom_of_le hrU) ⟨p, hpr⟩ _).trans e, diff --git a/src/algebraic_geometry/function_field.lean b/src/algebraic_geometry/function_field.lean index a56d5c1e1635c..6e8f90ac361b0 100644 --- a/src/algebraic_geometry/function_field.lean +++ b/src/algebraic_geometry/function_field.lean @@ -36,7 +36,7 @@ noncomputable abbreviation Scheme.germ_to_function_field [irreducible_space X.carrier] (U : opens X.carrier) [h : nonempty U] : X.presheaf.obj (op U) ⟶ X.function_field := X.presheaf.germ ⟨generic_point X.carrier, - ((generic_point_spec X.carrier).mem_open_set_iff U.prop).mpr (by simpa using h)⟩ + ((generic_point_spec X.carrier).mem_open_set_iff U.is_open).mpr (by simpa using h)⟩ noncomputable instance [irreducible_space X.carrier] (U : opens X.carrier) [nonempty U] : @@ -53,11 +53,11 @@ begin intro ha, replace ha := ne_of_apply_ne _ ha, have hs : generic_point X.carrier ∈ RingedSpace.basic_open _ s, - { rw [← opens.mem_coe, (generic_point_spec X.carrier).mem_open_set_iff, set.top_eq_univ, + { rw [← set_like.mem_coe, (generic_point_spec X.carrier).mem_open_set_iff, set.top_eq_univ, set.univ_inter, set.nonempty_iff_ne_empty, ne.def, ← opens.coe_bot, - subtype.coe_injective.eq_iff, ← opens.empty_eq], + ← set_like.ext'_iff], erw basic_open_eq_bot_iff, - exacts [ha, (RingedSpace.basic_open _ _).prop] }, + exacts [ha, (RingedSpace.basic_open _ _).is_open] }, have := (X.presheaf.germ ⟨_, hs⟩).is_unit_map (RingedSpace.is_unit_res_basic_open _ s), rwa Top.presheaf.germ_res_apply at this end @@ -145,7 +145,7 @@ end lemma is_affine_open.prime_ideal_of_generic_point {X : Scheme} [is_integral X] {U : opens X.carrier} (hU : is_affine_open U) [h : nonempty U] : hU.prime_ideal_of ⟨generic_point X.carrier, - ((generic_point_spec X.carrier).mem_open_set_iff U.prop).mpr (by simpa using h)⟩ = + ((generic_point_spec X.carrier).mem_open_set_iff U.is_open).mpr (by simpa using h)⟩ = generic_point (Scheme.Spec.obj $ op $ X.presheaf.obj $ op U).carrier := begin haveI : is_affine _ := hU, diff --git a/src/algebraic_geometry/locally_ringed_space.lean b/src/algebraic_geometry/locally_ringed_space.lean index 3e7c3c1bcbcb1..a1f0c6682c4a1 100644 --- a/src/algebraic_geometry/locally_ringed_space.lean +++ b/src/algebraic_geometry/locally_ringed_space.lean @@ -249,12 +249,9 @@ end @[simp] lemma basic_open_zero (X : LocallyRingedSpace) (U : opens X.carrier) : X.to_RingedSpace.basic_open (0 : X.presheaf.obj $ op U) = ⊥ := begin - ext, - simp only [set.mem_empty_iff_false, - topological_space.opens.mem_coe, opens.coe_bot, iff_false, RingedSpace.basic_open, - is_unit_zero_iff, set.mem_set_of_eq, map_zero], - rintro ⟨⟨y, _⟩, h, e⟩, - exact zero_ne_one' (X.presheaf.stalk y) h, + simp only [RingedSpace.basic_open, is_unit_zero_iff, map_zero, + zero_ne_one' (X.presheaf.stalk _), set.set_of_false, set.image_empty], + refl end instance component_nontrivial (X : LocallyRingedSpace) (U : opens X.carrier) diff --git a/src/algebraic_geometry/locally_ringed_space/has_colimits.lean b/src/algebraic_geometry/locally_ringed_space/has_colimits.lean index 15e89cc7ced7a..34efd4dff6154 100644 --- a/src/algebraic_geometry/locally_ringed_space/has_colimits.lean +++ b/src/algebraic_geometry/locally_ringed_space/has_colimits.lean @@ -199,11 +199,11 @@ begin have hV : (coequalizer.π f.1 g.1).base ⁻¹' ((coequalizer.π f.1 g.1).base '' V.1) = V.1 := image_basic_open_image_preimage f g U s, have hV' : V = ⟨(coequalizer.π f.1 g.1).base ⁻¹' - ((coequalizer.π f.1 g.1).base '' V.1), hV.symm ▸ V.2⟩ := subtype.eq hV.symm, - have V_open : is_open (((coequalizer.π f.val g.val).base) '' V.val) := + ((coequalizer.π f.1 g.1).base '' V.1), hV.symm ▸ V.2⟩ := set_like.ext' hV.symm, + have V_open : is_open (((coequalizer.π f.val g.val).base) '' V.1) := image_basic_open_image_open f g U s, have VleU : - (⟨((coequalizer.π f.val g.val).base) '' V.val, V_open⟩ : topological_space.opens _) ≤ U, + (⟨((coequalizer.π f.val g.val).base) '' V.1, V_open⟩ : topological_space.opens _) ≤ U, { exact set.image_subset_iff.mpr (Y.to_RingedSpace.basic_open_le _) }, have hxV : x ∈ V := ⟨⟨_, hU⟩, ha, rfl⟩, diff --git a/src/algebraic_geometry/morphisms/quasi_compact.lean b/src/algebraic_geometry/morphisms/quasi_compact.lean index 15b6850b92e16..b3281898a3059 100644 --- a/src/algebraic_geometry/morphisms/quasi_compact.lean +++ b/src/algebraic_geometry/morphisms/quasi_compact.lean @@ -72,27 +72,24 @@ lemma is_compact_open_iff_eq_finset_affine_union {X : Scheme} (U : set X.carrier is_compact U ∧ is_open U ↔ ∃ (s : set X.affine_opens), s.finite ∧ U = ⋃ (i : X.affine_opens) (h : i ∈ s), i := begin - apply opens.is_compact_open_iff_eq_finite_Union_of_is_basis + apply opens.is_basis.is_compact_open_iff_eq_finite_Union (coe : X.affine_opens → opens X.carrier), { rw subtype.range_coe, exact is_basis_affine_open X }, - { intro i, exact i.2.is_compact } + { exact λ i, i.2.is_compact } end lemma is_compact_open_iff_eq_basic_open_union {X : Scheme} [is_affine X] (U : set X.carrier) : is_compact U ∧ is_open U ↔ ∃ (s : set (X.presheaf.obj (op ⊤))), s.finite ∧ U = ⋃ (i : X.presheaf.obj (op ⊤)) (h : i ∈ s), X.basic_open i := -begin - apply opens.is_compact_open_iff_eq_finite_Union_of_is_basis, - { exact is_basis_basic_open X }, - { intro i, exact ((top_is_affine_open _).basic_open_is_affine _).is_compact } -end +(is_basis_basic_open X).is_compact_open_iff_eq_finite_Union _ + (λ i, ((top_is_affine_open _).basic_open_is_affine _).is_compact) _ lemma quasi_compact_iff_forall_affine : quasi_compact f ↔ ∀ U : opens Y.carrier, is_affine_open U → is_compact (f.1.base ⁻¹' (U : set Y.carrier)) := begin rw quasi_compact_iff, - refine ⟨λ H U hU, H U U.prop hU.is_compact, _⟩, + refine ⟨λ H U hU, H U U.is_open hU.is_compact, _⟩, intros H U hU hU', obtain ⟨S, hS, rfl⟩ := (is_compact_open_iff_eq_finset_affine_union U).mp ⟨hU', hU⟩, simp only [set.preimage_Union, subtype.val_eq_coe], @@ -123,7 +120,7 @@ lemma is_compact_basic_open (X : Scheme) {U : opens X.carrier} (hU : is_compact begin classical, refine ((is_compact_open_iff_eq_finset_affine_union _).mpr _).1, - obtain ⟨s, hs, e⟩ := (is_compact_open_iff_eq_finset_affine_union _).mp ⟨hU, U.prop⟩, + obtain ⟨s, hs, e⟩ := (is_compact_open_iff_eq_finset_affine_union _).mp ⟨hU, U.is_open⟩, let g : s → X.affine_opens, { intro V, use V.1 ⊓ X.basic_open f, @@ -135,7 +132,8 @@ begin exact is_affine_open.basic_open_is_affine V.1.prop _ }, haveI : finite s := hs.to_subtype, refine ⟨set.range g, set.finite_range g, _⟩, - refine (set.inter_eq_right_iff_subset.mpr (RingedSpace.basic_open_le _ _)).symm.trans _, + refine (set.inter_eq_right_iff_subset.mpr (set_like.coe_subset_coe.2 $ + RingedSpace.basic_open_le _ _)).symm.trans _, rw [e, set.Union₂_inter], apply le_antisymm; apply set.Union₂_subset, { intros i hi, diff --git a/src/algebraic_geometry/morphisms/quasi_separated.lean b/src/algebraic_geometry/morphisms/quasi_separated.lean index 5ccde009c11cf..8b75d385dcf82 100644 --- a/src/algebraic_geometry/morphisms/quasi_separated.lean +++ b/src/algebraic_geometry/morphisms/quasi_separated.lean @@ -262,7 +262,7 @@ lemma is_affine_open.is_quasi_separated {X : Scheme} {U : opens X.carrier} (hU : is_quasi_separated (U : set X.carrier) := begin rw is_quasi_separated_iff_quasi_separated_space, - exacts [@@algebraic_geometry.quasi_separated_space_of_is_affine _ hU, U.prop], + exacts [@@algebraic_geometry.quasi_separated_space_of_is_affine _ hU, U.is_open], end lemma quasi_separated_of_comp {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) diff --git a/src/algebraic_geometry/morphisms/ring_hom_properties.lean b/src/algebraic_geometry/morphisms/ring_hom_properties.lean index b9dfba0ca93c5..e389bdb008d2d 100644 --- a/src/algebraic_geometry/morphisms/ring_hom_properties.lean +++ b/src/algebraic_geometry/morphisms/ring_hom_properties.lean @@ -86,11 +86,10 @@ begin { apply ring_hom.to_algebra, refine X.presheaf.map (@hom_of_le _ _ ((is_open_map.functor _).obj _) ((is_open_map.functor _).obj _) _).op, - rw [opens.le_def], + rw [← set_like.coe_subset_coe], dsimp, - change coe '' (coe '' set.univ) ⊆ coe '' set.univ, - rw [subtype.coe_image_univ, subtype.coe_image_univ], - exact set.image_preimage_subset _ _ }, + simp only [set.image_univ, subtype.range_coe, set.image_subset_iff], + refl }, { exact algebraic_geometry.Γ_restrict_is_localization Y r }, { rw ← U.open_embedding_obj_top at hU, dsimp [Scheme.Γ_obj_op, Scheme.Γ_map_op, Scheme.restrict], diff --git a/src/algebraic_geometry/open_immersion.lean b/src/algebraic_geometry/open_immersion.lean index 92f56540532a7..f14801af53854 100644 --- a/src/algebraic_geometry/open_immersion.lean +++ b/src/algebraic_geometry/open_immersion.lean @@ -163,16 +163,15 @@ instance comp {Z : PresheafedSpace C} (f : X ⟶ Y) [hf : is_open_immersion f] ( apply_with is_iso.comp_is_iso { instances := ff }, swap, { have : (opens.map g.base).obj (h.functor.obj U) = hf.open_functor.obj U, - { dsimp only [opens.map, is_open_map.functor, PresheafedSpace.comp_base], - congr' 1, - rw [coe_comp, ←set.image_image, set.preimage_image_eq _ hg.base_open.inj] }, + { ext1, + dsimp only [opens.map_coe, is_open_map.functor_obj_coe, comp_base], + rw [coe_comp, ← set.image_image, set.preimage_image_eq _ hg.base_open.inj] }, rw this, apply_instance }, { have : h.functor.obj U = hg.open_functor.obj (hf.open_functor.obj U), - { dsimp only [is_open_map.functor], - congr' 1, - rw [comp_base, coe_comp, ←set.image_image], - congr }, + { ext1, + dsimp only [is_open_map.functor_obj_coe], + rw [comp_base, coe_comp, ←set.image_image] }, rw this, apply_instance } end } @@ -221,9 +220,9 @@ by { erw ← category.assoc, rw [is_iso.comp_inv_eq, f.c.naturality], congr } @[reassoc] lemma app_inv_app' (U : opens Y) (hU : (U : set Y) ⊆ set.range f.base) : f.c.app (op U) ≫ H.inv_app ((opens.map f.base).obj U) = Y.presheaf.map (eq_to_hom (by - { apply has_le.le.antisymm, + { apply le_antisymm, { exact set.image_preimage_subset f.base U.1 }, - { change U ⊆ _, + { rw [← set_like.coe_subset_coe], refine has_le.le.trans_eq _ (@set.image_preimage_eq_inter_range _ _ f.base U.1).symm, exact set.subset_inter_iff.mpr ⟨λ _ h, h, hU⟩ } })).op := by { erw ← category.assoc, rw [is_iso.comp_inv_eq, f.c.naturality], congr } @@ -244,11 +243,8 @@ instance of_restrict {X : Top} (Y : PresheafedSpace C) {f : X ⟶ Y.carrier} begin dsimp, have : (opens.map f).obj (hf.is_open_map.functor.obj U) = U, - { cases U, - dsimp only [opens.map, is_open_map.functor], - congr' 1, - rw set.preimage_image_eq _ hf.inj, - refl }, + { ext1, + exact set.preimage_image_eq _ hf.inj }, convert (show is_iso (Y.presheaf.map (𝟙 _)), from infer_instance), { apply subsingleton.helim, rw this }, @@ -376,7 +372,6 @@ def pullback_cone_of_left_lift : s.X ⟶ (pullback_cone_of_left f g).X := conv_lhs { erw ← this, rw coe_comp, erw ← set.preimage_preimage }, erw set.preimage_image_eq _ (Top.snd_open_embedding_of_left_open_embedding hf.base_open g.base).inj, - simp, end)), naturality' := λ U V i, begin @@ -810,7 +805,7 @@ end lemma image_preimage_is_empty (j : discrete ι) (h : i ≠ j) (U : opens (F.obj i)) : (opens.map (colimit.ι (F ⋙ SheafedSpace.forget_to_PresheafedSpace) j).base).obj ((opens.map (preserves_colimit_iso SheafedSpace.forget_to_PresheafedSpace F).inv.base).obj - ((sigma_ι_open_embedding F i).is_open_map.functor.obj U)) = ∅ := + ((sigma_ι_open_embedding F i).is_open_map.functor.obj U)) = ⊥ := begin ext, apply iff_false_intro, @@ -1595,7 +1590,7 @@ lemma range_pullback_to_base_of_left : set.range f.1.base ∩ set.range g.1.base := begin rw [pullback.condition, Scheme.comp_val_base, coe_comp, set.range_comp, - range_pullback_snd_of_left, opens.map_obj, subtype.coe_mk, set.image_preimage_eq_inter_range, + range_pullback_snd_of_left, opens.map_obj, opens.coe_mk, set.image_preimage_eq_inter_range, set.inter_comm], end @@ -1604,7 +1599,7 @@ lemma range_pullback_to_base_of_right : set.range g.1.base ∩ set.range f.1.base := begin rw [Scheme.comp_val_base, coe_comp, set.range_comp, range_pullback_fst_of_right, opens.map_obj, - subtype.coe_mk, set.image_preimage_eq_inter_range, set.inter_comm], + opens.coe_mk, set.image_preimage_eq_inter_range, set.inter_comm], end /-- @@ -1746,9 +1741,8 @@ lemma Scheme.restrict_functor_map_app_aux {U V : opens X.carrier} (i : U ⟶ V) ((opens.map (X.restrict_functor.map i).1.val.base).obj W) ≤ V.open_embedding.is_open_map.functor.obj W := begin - simp only [set.image_congr, subtype.mk_le_mk, is_open_map.functor, set.image_subset_iff, - Scheme.restrict_functor_map_base, opens.map, subtype.coe_mk, opens.inclusion_apply, - set.le_eq_subset], + simp only [← set_like.coe_subset_coe, is_open_map.functor_obj_coe, set.image_subset_iff, + Scheme.restrict_functor_map_base, opens.map_coe, opens.inclusion_apply], rintros _ h, exact ⟨_, h, rfl⟩, end diff --git a/src/algebraic_geometry/presheafed_space.lean b/src/algebraic_geometry/presheafed_space.lean index cc379a666307c..575dca4468f38 100644 --- a/src/algebraic_geometry/presheafed_space.lean +++ b/src/algebraic_geometry/presheafed_space.lean @@ -26,7 +26,7 @@ open category_theory.category category_theory.functor variables (C : Type u) [category.{v} C] -local attribute [tidy] tactic.op_induction' +local attribute [tidy] tactic.op_induction' tactic.auto_cases_opens namespace algebraic_geometry @@ -288,7 +288,7 @@ instance of_restrict_mono {U : Top} (X : PresheafedSpace C) (f : U ⟶ X.1) ext V, { induction V using opposite.rec, have hV : (opens.map (X.of_restrict hf).base).obj (hf.is_open_map.functor.obj V) = V, - { cases V, simp[opens.map, set.preimage_image_eq _ hf.inj] }, + { ext1, exact set.preimage_image_eq _ hf.inj }, haveI : is_iso (hf.is_open_map.adjunction.counit.app (unop (op (hf.is_open_map.functor.obj V)))) := (nat_iso.is_iso_app_of_is_iso (whisker_left diff --git a/src/algebraic_geometry/presheafed_space/gluing.lean b/src/algebraic_geometry/presheafed_space/gluing.lean index 144ed3c721d9e..df84093cc8460 100644 --- a/src/algebraic_geometry/presheafed_space/gluing.lean +++ b/src/algebraic_geometry/presheafed_space/gluing.lean @@ -205,8 +205,8 @@ lemma ι_image_preimage_eq (i j : D.J) (U : opens (D.U i).carrier) : (D.f_open j i).open_functor.obj ((opens.map (𝖣 .t j i).base).obj ((opens.map (𝖣 .f i j).base).obj U)) := begin - dsimp only [opens.map, is_open_map.functor], - congr' 1, + ext1, + dsimp only [opens.map_coe, is_open_map.functor_obj_coe], rw [← (show _ = (𝖣 .ι i).base, from 𝖣 .ι_glued_iso_inv (PresheafedSpace.forget _) i), ← (show _ = (𝖣 .ι j).base, from 𝖣 .ι_glued_iso_inv (PresheafedSpace.forget _) j), coe_comp, coe_comp, set.image_comp, set.preimage_comp, set.preimage_image_eq], @@ -284,8 +284,9 @@ begin rcases j with (⟨j, k⟩|j), { refine D.opens_image_preimage_map i j U ≫ (D.f j k).c.app _ ≫ (D.V (j, k)).presheaf.map (eq_to_hom _), - dsimp only [functor.op, opens.map, unop_op], - congr' 2, + rw [functor.op_obj], + congr' 1, ext1, + dsimp only [functor.op_obj, opens.map_coe, unop_op, is_open_map.functor_obj_coe], rw set.preimage_preimage, change (D.f j k ≫ 𝖣 .ι j).base ⁻¹' _ = _, congr' 3, diff --git a/src/algebraic_geometry/presheafed_space/has_colimits.lean b/src/algebraic_geometry/presheafed_space/has_colimits.lean index 11619962ca2ac..eefd2a12e5c49 100644 --- a/src/algebraic_geometry/presheafed_space/has_colimits.lean +++ b/src/algebraic_geometry/presheafed_space/has_colimits.lean @@ -52,6 +52,7 @@ namespace algebraic_geometry namespace PresheafedSpace local attribute [simp] eq_to_hom_map +local attribute [tidy] tactic.auto_cases_opens @[simp] lemma map_id_c_app (F : J ⥤ PresheafedSpace.{v} C) (j) (U) : @@ -359,9 +360,8 @@ begin fapply nat_iso.of_components, { intro X, refine ((F.obj (unop X)).presheaf.map_iso (eq_to_iso _)), - dsimp only [functor.op, unop_op, opens.map], - congr' 2, - rw set.preimage_preimage, + simp only [functor.op_obj, unop_op, op_inj_iff, opens.map_coe, set_like.ext'_iff, + set.preimage_preimage], simp_rw ← comp_app, congr' 2, exact ι_preserves_colimits_iso_inv (forget C) F (unop X) }, diff --git a/src/algebraic_geometry/prime_spectrum/basic.lean b/src/algebraic_geometry/prime_spectrum/basic.lean index 711ffcd6a0da6..ad98e61dcd8fd 100644 --- a/src/algebraic_geometry/prime_spectrum/basic.lean +++ b/src/algebraic_geometry/prime_spectrum/basic.lean @@ -672,14 +672,14 @@ section basic_open /-- `basic_open r` is the open subset containing all prime ideals not containing `r`. -/ def basic_open (r : R) : topological_space.opens (prime_spectrum R) := -{ val := { x | r ∉ x.as_ideal }, - property := ⟨{r}, set.ext $ λ x, set.singleton_subset_iff.trans $ not_not.symm⟩ } +{ carrier := { x | r ∉ x.as_ideal }, + is_open' := ⟨{r}, set.ext $ λ x, set.singleton_subset_iff.trans $ not_not.symm⟩ } @[simp] lemma mem_basic_open (f : R) (x : prime_spectrum R) : x ∈ basic_open f ↔ f ∉ x.as_ideal := iff.rfl lemma is_open_basic_open {a : R} : is_open ((basic_open a) : set (prime_spectrum R)) := -(basic_open a).property +(basic_open a).is_open @[simp] lemma basic_open_eq_zero_locus_compl (r : R) : (basic_open r : set (prime_spectrum R)) = (zero_locus {r})ᶜ := @@ -693,8 +693,8 @@ topological_space.opens.ext $ by simp lemma basic_open_le_basic_open_iff (f g : R) : basic_open f ≤ basic_open g ↔ f ∈ (ideal.span ({g} : set R)).radical := -by rw [topological_space.opens.le_def, basic_open_eq_zero_locus_compl, - basic_open_eq_zero_locus_compl, set.le_eq_subset, set.compl_subset_compl, +by rw [← set_like.coe_subset_coe, basic_open_eq_zero_locus_compl, + basic_open_eq_zero_locus_compl, set.compl_subset_compl, zero_locus_subset_zero_locus_singleton_iff] lemma basic_open_mul (f g : R) : basic_open (f * g) = basic_open f ⊓ basic_open g := @@ -758,7 +758,7 @@ end lemma basic_open_eq_bot_iff (f : R) : basic_open f = ⊥ ↔ is_nilpotent f := begin - rw [← subtype.coe_injective.eq_iff, basic_open_eq_zero_locus_compl], + rw [← topological_space.opens.coe_inj, basic_open_eq_zero_locus_compl], simp only [set.eq_univ_iff_forall, set.singleton_subset_iff, topological_space.opens.coe_bot, nilpotent_iff_mem_prime, set.compl_empty_iff, mem_zero_locus, set_like.mem_coe], diff --git a/src/algebraic_geometry/projective_spectrum/scheme.lean b/src/algebraic_geometry/projective_spectrum/scheme.lean index 78fb0a006f8d7..4bfbcf04dcd7e 100644 --- a/src/algebraic_geometry/projective_spectrum/scheme.lean +++ b/src/algebraic_geometry/projective_spectrum/scheme.lean @@ -253,7 +253,7 @@ begin classical, ext1 y, split; intros hy, { refine ⟨y.2, _⟩, - rw [set.mem_preimage, opens.mem_coe, prime_spectrum.mem_basic_open] at hy, + rw [set.mem_preimage, set_like.mem_coe, prime_spectrum.mem_basic_open] at hy, rw projective_spectrum.mem_coe_basic_open, intro a_mem_y, apply hy, @@ -265,7 +265,7 @@ begin { change y.1 ∈ _ at hy, rcases hy with ⟨hy1, hy2⟩, rw projective_spectrum.mem_coe_basic_open at hy1 hy2, - rw [set.mem_preimage, to_fun, opens.mem_coe, prime_spectrum.mem_basic_open], + rw [set.mem_preimage, to_fun, set_like.mem_coe, prime_spectrum.mem_basic_open], intro rid, dsimp at rid, rcases mem_carrier.clear_denominator 𝒜 _ rid with ⟨c, N, acd, eq1⟩, rw [algebra.smul_def] at eq1, diff --git a/src/algebraic_geometry/projective_spectrum/topology.lean b/src/algebraic_geometry/projective_spectrum/topology.lean index dfd2ee4a7847e..cd87b1507213d 100644 --- a/src/algebraic_geometry/projective_spectrum/topology.lean +++ b/src/algebraic_geometry/projective_spectrum/topology.lean @@ -327,8 +327,8 @@ section basic_open /-- `basic_open r` is the open subset containing all prime ideals not containing `r`. -/ def basic_open (r : A) : topological_space.opens (projective_spectrum 𝒜) := -{ val := { x | r ∉ x.as_homogeneous_ideal }, - property := ⟨{r}, set.ext $ λ x, set.singleton_subset_iff.trans $ not_not.symm⟩ } +{ carrier := { x | r ∉ x.as_homogeneous_ideal }, + is_open' := ⟨{r}, set.ext $ λ x, set.singleton_subset_iff.trans $ not_not.symm⟩ } @[simp] lemma mem_basic_open (f : A) (x : projective_spectrum 𝒜) : x ∈ basic_open 𝒜 f ↔ f ∉ x.as_homogeneous_ideal := iff.rfl @@ -338,7 +338,7 @@ lemma mem_coe_basic_open (f : A) (x : projective_spectrum 𝒜) : lemma is_open_basic_open {a : A} : is_open ((basic_open 𝒜 a) : set (projective_spectrum 𝒜)) := -(basic_open 𝒜 a).property +(basic_open 𝒜 a).is_open @[simp] lemma basic_open_eq_zero_locus_compl (r : A) : (basic_open 𝒜 r : set (projective_spectrum 𝒜)) = (zero_locus 𝒜 {r})ᶜ := diff --git a/src/algebraic_geometry/properties.lean b/src/algebraic_geometry/properties.lean index 52d00b6bfb420..a67a35022dfcf 100644 --- a/src/algebraic_geometry/properties.lean +++ b/src/algebraic_geometry/properties.lean @@ -142,7 +142,8 @@ begin intros X U, apply h₁, intro x, - obtain ⟨_,⟨j,rfl⟩,hx,i⟩ := X.affine_basis_cover_is_basis.exists_subset_of_mem_open x.prop U.2, + obtain ⟨_, ⟨j, rfl⟩, hx, i⟩ := X.affine_basis_cover_is_basis.exists_subset_of_mem_open + (set_like.mem_coe.2 x.prop) U.is_open, let U' : opens _ := ⟨_, (X.affine_basis_cover.is_open j).base_open.open_range⟩, let i' : U' ⟶ U := hom_of_le i, @@ -151,7 +152,7 @@ begin apply h₂', apply h₃ end -. + lemma reduce_to_affine_nbhd (P : ∀ (X : Scheme) (x : X.carrier), Prop) (h₁ : ∀ (R : CommRing) (x : prime_spectrum R), P (Scheme.Spec.obj $ op R) x) (h₂ : ∀ {X Y} (f : X ⟶ Y) [is_open_immersion f] (x : X.carrier), P X x → P Y (f.1.base x)) : @@ -176,17 +177,17 @@ begin obtain ⟨V, hx, i, H⟩ := hx x, unfreezingI { specialize H (X.presheaf.map i.op s) }, erw Scheme.basic_open_res at H, - rw [hs, ← subtype.coe_injective.eq_iff, inf_bot_eq] at H, - specialize H rfl ⟨x, hx⟩, + rw [hs] at H, + specialize H inf_bot_eq ⟨x, hx⟩, erw Top.presheaf.germ_res_apply at H, exact H }, { rintros X Y f hf, - have e : (f.val.base) ⁻¹' set.range ⇑(f.val.base) = ⊤, - { rw [← set.image_univ, set.preimage_image_eq _ hf.base_open.inj, set.top_eq_univ] }, + have e : (f.val.base) ⁻¹' set.range ⇑(f.val.base) = set.univ, + { rw [← set.image_univ, set.preimage_image_eq _ hf.base_open.inj] }, refine ⟨_, _, e, rfl, _⟩, rintros H hX s hs ⟨_, x, rfl⟩, unfreezingI { haveI := is_reduced_of_open_immersion f }, - specialize H (f.1.c.app _ s) _ ⟨x, by { change x ∈ (f.val.base) ⁻¹' _, rw e, trivial }⟩, + specialize H (f.1.c.app _ s) _ ⟨x, by { rw [opens.mem_mk, e], trivial }⟩, { rw [← Scheme.preimage_basic_open, hs], ext1, simp [opens.map] }, { erw ← PresheafedSpace.stalk_map_germ_apply f.1 ⟨_,_⟩ ⟨x,_⟩ at H, apply_fun (inv $ PresheafedSpace.stalk_map f.val x) at H, @@ -229,8 +230,8 @@ instance is_reduced_of_is_integral [is_integral X] : is_reduced X := begin constructor, intro U, - cases U.1.eq_empty_or_nonempty, - { have : U = ∅ := subtype.eq h, + cases U.1.eq_empty_or_nonempty with h h, + { have : U = ⊥ := set_like.ext' h, haveI := CommRing.subsingleton_of_is_terminal (X.sheaf.is_terminal_of_eq_empty this), change _root_.is_reduced (X.sheaf.val.obj (op U)), apply_instance }, @@ -339,7 +340,7 @@ begin revert hx, contrapose!, simp_rw [← opens.not_nonempty_iff_eq_bot, not_not], - apply nonempty_preirreducible_inter U.prop (RingedSpace.basic_open _ _).prop, + apply nonempty_preirreducible_inter U.is_open (RingedSpace.basic_open _ _).is_open, simpa using H end diff --git a/src/algebraic_geometry/ringed_space.lean b/src/algebraic_geometry/ringed_space.lean index 506bc9d9cdbb6..9018859254f69 100644 --- a/src/algebraic_geometry/ringed_space.lean +++ b/src/algebraic_geometry/ringed_space.lean @@ -68,7 +68,7 @@ begin choose V iVU m h_unit using λ x : U, X.is_unit_res_of_is_unit_germ U f x (h x), have hcover : U ≤ supr V, { intros x hxU, - rw [opens.mem_coe, opens.mem_supr], + rw [opens.mem_supr], exact ⟨⟨x, hxU⟩, m ⟨x, hxU⟩⟩ }, -- Let `g x` denote the inverse of `f` in `U x`. choose g hg using λ x : U, is_unit.exists_right_inv (h_unit x), @@ -97,8 +97,8 @@ The basic open of a section `f` is the set of all points `x`, such that the germ `x` is a unit. -/ def basic_open {U : opens X} (f : X.presheaf.obj (op U)) : opens X := -{ val := coe '' { x : U | is_unit (X.presheaf.germ x f) }, - property := begin +{ carrier := coe '' { x : U | is_unit (X.presheaf.germ x f) }, + is_open' := begin rw is_open_iff_forall_mem_open, rintros _ ⟨x, hx, rfl⟩, obtain ⟨V, i, hxV, hf⟩ := X.is_unit_res_of_is_unit_germ U f x hx, diff --git a/src/algebraic_geometry/stalks.lean b/src/algebraic_geometry/stalks.lean index ce1bcf0f929db..a702694af2196 100644 --- a/src/algebraic_geometry/stalks.lean +++ b/src/algebraic_geometry/stalks.lean @@ -27,7 +27,7 @@ open opposite variables {C : Type u} [category.{v} C] [has_colimits C] -local attribute [tidy] tactic.op_induction' +local attribute [tidy] tactic.op_induction' tactic.auto_cases_opens open Top.presheaf diff --git a/src/algebraic_geometry/structure_sheaf.lean b/src/algebraic_geometry/structure_sheaf.lean index c9dd40eec9e22..a557011cfabb5 100644 --- a/src/algebraic_geometry/structure_sheaf.lean +++ b/src/algebraic_geometry/structure_sheaf.lean @@ -608,9 +608,9 @@ lemma locally_const_basic_open (U : opens (prime_spectrum.Top R)) begin -- First, any section `s` can be represented as a fraction `f/g` on some open neighborhood of `x` -- and we may pass to a `basic_open h`, since these form a basis - obtain ⟨V, (hxV : x.1 ∈ V.1), iVU, f, g, (hVDg : V ⊆ basic_open g), s_eq⟩ := + obtain ⟨V, (hxV : x.1 ∈ V.1), iVU, f, g, (hVDg : V ≤ basic_open g), s_eq⟩ := exists_const R U s x.1 x.2, - obtain ⟨_, ⟨h, rfl⟩, hxDh, (hDhV : basic_open h ⊆ V)⟩ := + obtain ⟨_, ⟨h, rfl⟩, hxDh, (hDhV : basic_open h ≤ V)⟩ := is_topological_basis_basic_opens.exists_subset_of_mem_open hxV V.2, -- The problem is of course, that `g` and `h` don't need to coincide. -- But, since `basic_open h ≤ basic_open g`, some power of `h` must be a multiple of `g` @@ -644,11 +644,11 @@ A local representation of a section `s` as fractions `a i / h i` on finitely man -/ lemma normalize_finite_fraction_representation (U : opens (prime_spectrum.Top R)) (s : (structure_sheaf R).1.obj (op U)) {ι : Type*} (t : finset ι) (a h : ι → R) - (iDh : Π i : ι, basic_open (h i) ⟶ U) (h_cover : U.1 ⊆ ⋃ i ∈ t, (basic_open (h i)).1) + (iDh : Π i : ι, basic_open (h i) ⟶ U) (h_cover : U ≤ ⨆ i ∈ t, basic_open (h i)) (hs : ∀ i : ι, const R (a i) (h i) (basic_open (h i)) (λ y hy, hy) = (structure_sheaf R).1.map (iDh i).op s) : ∃ (a' h' : ι → R) (iDh' : Π i : ι, (basic_open (h' i)) ⟶ U), - (U.1 ⊆ ⋃ i ∈ t, (basic_open (h' i)).1) ∧ + (U ≤ ⨆ i ∈ t, basic_open (h' i)) ∧ (∀ i j ∈ t, a' i * h' j = h' i * a' j) ∧ (∀ i ∈ t, (structure_sheaf R).1.map (iDh' i).op s = const R (a' i) (h' i) (basic_open (h' i)) (λ y hy, hy)) := @@ -735,19 +735,21 @@ begin choose a' h' iDh' hxDh' s_eq' using locally_const_basic_open R (basic_open f) s, -- Since basic opens are compact, we can pass to a finite subcover obtain ⟨t, ht_cover'⟩ := (is_compact_basic_open f).elim_finite_subcover - (λ (i : ι), (basic_open (h' i)).1) (λ i, is_open_basic_open) (λ x hx, _), + (λ (i : ι), basic_open (h' i)) (λ i, is_open_basic_open) (λ x hx, _), swap, { -- Here, we need to show that our basic opens actually form a cover of `basic_open f` rw set.mem_Union, exact ⟨⟨x,hx⟩, hxDh' ⟨x, hx⟩⟩ }, + simp only [← opens.coe_supr, set_like.coe_subset_coe] at ht_cover', -- We use the normalization lemma from above to obtain the relation `a i * h j = h i * a j` obtain ⟨a, h, iDh, ht_cover, ah_ha, s_eq⟩ := normalize_finite_fraction_representation R (basic_open f) s t a' h' iDh' ht_cover' s_eq', clear s_eq' iDh' hxDh' ht_cover' a' h', + simp only [← set_like.coe_subset_coe, opens.coe_supr] at ht_cover, -- Next we show that some power of `f` is a linear combination of the `h i` obtain ⟨n, hn⟩ : f ∈ (ideal.span (h '' ↑t)).radical, { rw [← vanishing_ideal_zero_locus_eq_radical, zero_locus_span], - simp_rw [subtype.val_eq_coe, basic_open_eq_zero_locus_compl] at ht_cover, + simp only [basic_open_eq_zero_locus_compl] at ht_cover, rw set.compl_subset_comm at ht_cover, -- Why doesn't `simp_rw` do this? simp_rw [set.compl_Union, compl_compl, ← zero_locus_Union, ← finset.set_bUnion_coe, ← set.image_eq_Union ] at ht_cover, diff --git a/src/measure_theory/measure/content.lean b/src/measure_theory/measure/content.lean index ed2cca7809d60..3ecbc9123f914 100644 --- a/src/measure_theory/measure/content.lean +++ b/src/measure_theory/measure/content.lean @@ -120,8 +120,8 @@ lemma inner_content_of_is_compact {K : set G} (h1K : is_compact K) (h2K : is_ope le_antisymm (supr₂_le $ λ K' hK', μ.mono _ ⟨K, h1K⟩ hK') (μ.le_inner_content _ _ subset.rfl) -lemma inner_content_empty : - μ.inner_content ∅ = 0 := +lemma inner_content_bot : + μ.inner_content ⊥ = 0 := begin refine le_antisymm _ (zero_le _), rw ←μ.empty, refine supr₂_le (λ K hK, _), @@ -157,9 +157,9 @@ begin { intros n s hn ih, rw [finset.sup_insert, finset.sum_insert hn], exact le_trans (μ.sup_le _ _) (add_le_add_left ih _) }}, refine supr₂_le (λ K hK, _), - obtain ⟨t, ht⟩ := K.is_compact.elim_finite_subcover _ (λ i, (U i).prop) _, swap, - { convert hK, rw [opens.supr_def, subtype.coe_mk] }, - rcases K.is_compact.finite_compact_cover t (coe ∘ U) (λ i _, (U _).prop) (by simp only [ht]) + obtain ⟨t, ht⟩ := K.is_compact.elim_finite_subcover _ (λ i, (U i).is_open) _, swap, + { rwa [← opens.coe_supr] }, + rcases K.is_compact.finite_compact_cover t (coe ∘ U) (λ i _, (U _).is_open) (by simp only [ht]) with ⟨K', h1K', h2K', h3K'⟩, let L : ℕ → compacts G := λ n, ⟨K' n, h1K' n⟩, convert le_trans (h3 t L) _, @@ -198,14 +198,14 @@ lemma inner_content_pos_of_is_mul_left_invariant [t2_space G] [group G] [topolog (K : compacts G) (hK : μ K ≠ 0) (U : opens G) (hU : (U : set G).nonempty) : 0 < μ.inner_content U := begin - have : (interior (U : set G)).nonempty, rwa [U.prop.interior_eq], + have : (interior (U : set G)).nonempty, rwa [U.is_open.interior_eq], rcases compact_covered_by_mul_left_translates K.2 this with ⟨s, hs⟩, suffices : μ K ≤ s.card * μ.inner_content U, { exact (ennreal.mul_pos_iff.mp $ hK.bot_lt.trans_le this).2 }, have : (K : set G) ⊆ ↑⨆ (g ∈ s), opens.comap (homeomorph.mul_left g).to_continuous_map U, { simpa only [opens.supr_def, opens.coe_comap, subtype.coe_mk] }, refine (μ.le_inner_content _ _ this).trans _, - refine (rel_supr_sum (μ.inner_content) (μ.inner_content_empty) (≤) + refine (rel_supr_sum (μ.inner_content) (μ.inner_content_bot) (≤) (μ.inner_content_Sup_nat) _ _).trans _, simp only [μ.is_mul_left_invariant_inner_content h3, finset.sum_const, nsmul_eq_mul, le_refl] end @@ -218,7 +218,7 @@ section outer_measure /-- Extending a content on compact sets to an outer measure on all sets. -/ protected def outer_measure : outer_measure G := -induced_outer_measure (λ U hU, μ.inner_content ⟨U, hU⟩) is_open_empty μ.inner_content_empty +induced_outer_measure (λ U hU, μ.inner_content ⟨U, hU⟩) is_open_empty μ.inner_content_bot variables [t2_space G] @@ -293,7 +293,7 @@ lemma outer_measure_caratheodory (A : set G) : measurable_set[μ.outer_measure.caratheodory] A ↔ ∀ (U : opens G), μ.outer_measure (U ∩ A) + μ.outer_measure (U \ A) ≤ μ.outer_measure U := begin - dsimp [opens], rw subtype.forall, + rw [opens.forall], apply induced_outer_measure_caratheodory, apply inner_content_Union_nat, apply inner_content_mono' @@ -318,7 +318,7 @@ begin intros U hU, rw μ.outer_measure_caratheodory, intro U', - rw μ.outer_measure_of_is_open ((U' : set G) ∩ U) (is_open.inter U'.prop hU), + rw μ.outer_measure_of_is_open ((U' : set G) ∩ U) (U'.is_open.inter hU), simp only [inner_content, supr_subtype'], rw [opens.coe_mk], haveI : nonempty {L : compacts G // (L : set G) ⊆ U' ∩ U} := ⟨⟨⊥, empty_subset _⟩⟩, rw [ennreal.supr_add], diff --git a/src/measure_theory/measure/haar.lean b/src/measure_theory/measure/haar.lean index 340eb715e8c07..4e38aba3cf793 100644 --- a/src/measure_theory/measure/haar.lean +++ b/src/measure_theory/measure/haar.lean @@ -322,10 +322,10 @@ begin have : is_compact (haar_product (K₀ : set G)), { apply is_compact_univ_pi, intro K, apply is_compact_Icc }, refine this.inter_Inter_nonempty (cl_prehaar K₀) (λ s, is_closed_closure) (λ t, _), - let V₀ := ⋂ (V ∈ t), (V : open_nhds_of 1).1, + let V₀ := ⋂ (V ∈ t), (V : open_nhds_of 1).carrier, have h1V₀ : is_open V₀, - { apply is_open_bInter, apply finset.finite_to_set, rintro ⟨V, hV⟩ h2V, exact hV.1 }, - have h2V₀ : (1 : G) ∈ V₀, { simp only [mem_Inter], rintro ⟨V, hV⟩ h2V, exact hV.2 }, + { apply is_open_bInter, apply finset.finite_to_set, rintro ⟨⟨V, hV₁⟩, hV₂⟩ h2V, exact hV₁ }, + have h2V₀ : (1 : G) ∈ V₀, { simp only [mem_Inter], rintro ⟨⟨V, hV₁⟩, hV₂⟩ h2V, exact hV₂ }, refine ⟨prehaar K₀ V₀, _⟩, split, { apply prehaar_mem_haar_product K₀, use 1, rwa h1V₀.interior_eq }, @@ -368,7 +368,7 @@ begin let eval : (compacts G → ℝ) → ℝ := λ f, f ⊥, have : continuous eval := continuous_apply ⊥, show chaar K₀ ∈ eval ⁻¹' {(0 : ℝ)}, - apply mem_of_subset_of_mem _ (chaar_mem_cl_prehaar K₀ ⟨set.univ, is_open_univ, mem_univ _⟩), + apply mem_of_subset_of_mem _ (chaar_mem_cl_prehaar K₀ ⊤), unfold cl_prehaar, rw is_closed.closure_subset_iff, { rintro _ ⟨U, ⟨h1U, h2U, h3U⟩, rfl⟩, apply prehaar_empty }, { apply continuous_iff_is_closed.mp this, exact is_closed_singleton }, @@ -380,7 +380,7 @@ begin let eval : (compacts G → ℝ) → ℝ := λ f, f K₀.to_compacts, have : continuous eval := continuous_apply _, show chaar K₀ ∈ eval ⁻¹' {(1 : ℝ)}, - apply mem_of_subset_of_mem _ (chaar_mem_cl_prehaar K₀ ⟨set.univ, is_open_univ, mem_univ _⟩), + apply mem_of_subset_of_mem _ (chaar_mem_cl_prehaar K₀ ⊤), unfold cl_prehaar, rw is_closed.closure_subset_iff, { rintro _ ⟨U, ⟨h1U, h2U, h3U⟩, rfl⟩, apply prehaar_self, rw h2U.interior_eq, exact ⟨1, h3U⟩ }, @@ -394,7 +394,7 @@ begin let eval : (compacts G → ℝ) → ℝ := λ f, f K₂ - f K₁, have : continuous eval := (continuous_apply K₂).sub (continuous_apply K₁), rw [← sub_nonneg], show chaar K₀ ∈ eval ⁻¹' (Ici (0 : ℝ)), - apply mem_of_subset_of_mem _ (chaar_mem_cl_prehaar K₀ ⟨set.univ, is_open_univ, mem_univ _⟩), + apply mem_of_subset_of_mem _ (chaar_mem_cl_prehaar K₀ ⊤), unfold cl_prehaar, rw is_closed.closure_subset_iff, { rintro _ ⟨U, ⟨h1U, h2U, h3U⟩, rfl⟩, simp only [mem_preimage, mem_Ici, eval, sub_nonneg], apply prehaar_mono _ h, rw h2U.interior_eq, exact ⟨1, h3U⟩ }, @@ -410,7 +410,7 @@ begin ((@continuous_add ℝ _ _ _).comp ((continuous_apply K₁).prod_mk (continuous_apply K₂))).sub (continuous_apply (K₁ ⊔ K₂)), rw [← sub_nonneg], show chaar K₀ ∈ eval ⁻¹' (Ici (0 : ℝ)), - apply mem_of_subset_of_mem _ (chaar_mem_cl_prehaar K₀ ⟨set.univ, is_open_univ, mem_univ _⟩), + apply mem_of_subset_of_mem _ (chaar_mem_cl_prehaar K₀ ⊤), unfold cl_prehaar, rw is_closed.closure_subset_iff, { rintro _ ⟨U, ⟨h1U, h2U, h3U⟩, rfl⟩, simp only [mem_preimage, mem_Ici, eval, sub_nonneg], apply prehaar_sup_le, rw h2U.interior_eq, exact ⟨1, h3U⟩ }, @@ -435,7 +435,7 @@ begin rw [eq_comm, ← sub_eq_zero], show chaar K₀ ∈ eval ⁻¹' {(0 : ℝ)}, let V := V₁ ∩ V₂, apply mem_of_subset_of_mem _ (chaar_mem_cl_prehaar K₀ - ⟨V⁻¹, (is_open.inter h2V₁ h2V₂).preimage continuous_inv, + ⟨⟨V⁻¹, (h2V₁.inter h2V₂).preimage continuous_inv⟩, by simp only [mem_inv, inv_one, h3V₁, h3V₂, V, mem_inter_iff, true_and]⟩), unfold cl_prehaar, rw is_closed.closure_subset_iff, { rintro _ ⟨U, ⟨h1U, h2U, h3U⟩, rfl⟩, @@ -457,7 +457,7 @@ begin let eval : (compacts G → ℝ) → ℝ := λ f, f (K.map _ $ continuous_mul_left g) - f K, have : continuous eval := (continuous_apply (K.map _ _)).sub (continuous_apply K), rw [← sub_eq_zero], show chaar K₀ ∈ eval ⁻¹' {(0 : ℝ)}, - apply mem_of_subset_of_mem _ (chaar_mem_cl_prehaar K₀ ⟨set.univ, is_open_univ, mem_univ _⟩), + apply mem_of_subset_of_mem _ (chaar_mem_cl_prehaar K₀ ⊤), unfold cl_prehaar, rw is_closed.closure_subset_iff, { rintro _ ⟨U, ⟨h1U, h2U, h3U⟩, rfl⟩, simp only [mem_singleton_iff, mem_preimage, eval, sub_eq_zero], diff --git a/src/tactic/auto_cases.lean b/src/tactic/auto_cases.lean index 04801a9ca70a8..95adf9b9aa6b3 100644 --- a/src/tactic/auto_cases.lean +++ b/src/tactic/auto_cases.lean @@ -49,9 +49,10 @@ meta def find_tac : expr → option auto_cases_tac end auto_cases /-- Applies `cases` or `induction` on the local_hypothesis `hyp : expr`. -/ -meta def auto_cases_at (hyp : expr) : tactic string := +meta def auto_cases_at (find : expr → option auto_cases.auto_cases_tac) (hyp : expr) : + tactic string := do t ← infer_type hyp >>= whnf, - match auto_cases.find_tac t with + match find t with | some atac := do atac.tac hyp, pp ← pp hyp, @@ -61,9 +62,9 @@ do t ← infer_type hyp >>= whnf, /-- Applies `cases` or `induction` on certain hypotheses. -/ @[hint_tactic] -meta def auto_cases : tactic string := +meta def auto_cases (find := tactic.auto_cases.find_tac) : tactic string := do l ← local_context, - results ← successes $ l.reverse.map auto_cases_at, + results ← successes $ l.reverse.map (auto_cases_at find), when (results.empty) $ fail "`auto_cases` did not find any hypotheses to apply `cases` or `induction` to", return (string.intercalate ", " results) diff --git a/src/topology/category/Top/opens.lean b/src/topology/category/Top/opens.lean index 251221c85dcbe..e69e5a9f4e7e5 100644 --- a/src/topology/category/Top/opens.lean +++ b/src/topology/category/Top/opens.lean @@ -98,7 +98,7 @@ The functor from open sets in `X` to `Top`, realising each open set as a topological space itself. -/ def to_Top (X : Top.{u}) : opens X ⥤ Top := -{ obj := λ U, ⟨U.val, infer_instance⟩, +{ obj := λ U, ⟨U, infer_instance⟩, map := λ U V i, ⟨λ x, ⟨x.1, i.le x.2⟩, (embedding.continuous_iff embedding_subtype_coe).2 continuous_induced_dom⟩ } @@ -110,7 +110,7 @@ rfl /-- The inclusion map from an open subset to the whole space, as a morphism in `Top`. -/ -@[simps] +@[simps { fully_applied := ff }] def inclusion {X : Top.{u}} (U : opens X) : (to_Top X).obj U ⟶ X := { to_fun := _, continuous_to_fun := continuous_subtype_coe } @@ -128,7 +128,7 @@ def inclusion_top_iso (X : Top.{u}) : (to_Top X).obj ⊤ ≅ X := /-- `opens.map f` gives the functor from open sets in Y to open set in X, given by taking preimages under f. -/ def map (f : X ⟶ Y) : opens Y ⥤ opens X := -{ obj := λ U, ⟨ f ⁻¹' U.val, U.property.preimage f.continuous ⟩, +{ obj := λ U, ⟨ f ⁻¹' U, U.is_open.preimage f.continuous ⟩, map := λ U V i, ⟨ ⟨ λ x h, i.le h ⟩ ⟩ }. lemma map_coe (f : X ⟶ Y) (U : opens Y) : @@ -178,7 +178,7 @@ rfl lemma map_supr (f : X ⟶ Y) {ι : Type*} (U : ι → opens Y) : (map f).obj (supr U) = supr ((map f).obj ∘ U) := begin - apply subtype.eq, rw [supr_def, supr_def, map_obj], + ext1, rw [supr_def, supr_def, map_obj], dsimp, rw set.preimage_Union, refl, end @@ -235,7 +235,10 @@ rfl eq_to_hom (congr_fun (congr_arg functor.obj (congr_arg map h.symm)) U) := rfl -/-- A homeomorphism of spaces gives an equivalence of categories of open sets. -/ +/-- A homeomorphism of spaces gives an equivalence of categories of open sets. + +TODO: define `order_iso.equivalence`, use it. +-/ @[simps] def map_map_iso {X Y : Top.{u}} (H : X ≅ Y) : opens Y ≌ opens X := { functor := map H.hom, inverse := map H.inv, diff --git a/src/topology/continuous_function/ideals.lean b/src/topology/continuous_function/ideals.lean index b51e302e6c6f7..22779d09edb55 100644 --- a/src/topology/continuous_function/ideals.lean +++ b/src/topology/continuous_function/ideals.lean @@ -315,7 +315,7 @@ variable (X) galois_insertion (opens_of_ideal : ideal C(X, 𝕜) → opens X) (λ s, ideal_of_set 𝕜 s) := { choice := λ I hI, opens_of_ideal I.closure, gc := λ I s, ideal_gc X 𝕜 I s, - le_l_u := λ s, (set_of_ideal_of_set_of_is_open 𝕜 s.prop).ge, + le_l_u := λ s, (set_of_ideal_of_set_of_is_open 𝕜 s.is_open).ge, choice_eq := λ I hI, congr_arg _ $ ideal.ext (set.ext_iff.mp (is_closed_of_closure_subset $ (ideal_of_set_of_ideal_eq_closure I ▸ hI : I.closure ≤ I)).closure_eq) } diff --git a/src/topology/continuous_function/t0_sierpinski.lean b/src/topology/continuous_function/t0_sierpinski.lean index a41e477f4054a..0c511fb9044e3 100644 --- a/src/topology/continuous_function/t0_sierpinski.lean +++ b/src/topology/continuous_function/t0_sierpinski.lean @@ -43,7 +43,7 @@ open subset `u` of `X`). The `u` coordinate of `product_of_mem_opens x` is given -/ def product_of_mem_opens : C(X, opens X → Prop) := { to_fun := λ x u, x ∈ u, - continuous_to_fun := continuous_pi_iff.2 (λ u, continuous_Prop.2 u.property) } + continuous_to_fun := continuous_pi_iff.2 (λ u, continuous_Prop.2 u.is_open) } lemma product_of_mem_opens_inducing : inducing (product_of_mem_opens X) := begin diff --git a/src/topology/gluing.lean b/src/topology/gluing.lean index 4c198b3bed6be..b254ae5a0ba85 100644 --- a/src/topology/gluing.lean +++ b/src/topology/gluing.lean @@ -259,7 +259,7 @@ begin rw preimage_image_eq_image, apply (D.f_open _ _).is_open_map, apply (D.t j i ≫ D.f i j).continuous_to_fun.is_open_preimage, - exact U.property + exact U.is_open end lemma ι_open_embedding (i : D.J) : open_embedding (𝖣 .ι i) := diff --git a/src/topology/sets/opens.lean b/src/topology/sets/opens.lean index 527ff91f7fcf2..a64f094a3f5b2 100644 --- a/src/topology/sets/opens.lean +++ b/src/topology/sets/opens.lean @@ -8,6 +8,7 @@ import topology.bases import topology.homeomorph import topology.continuous_function.basic import order.compactly_generated +import tactic.auto_cases /-! # Open sets @@ -18,57 +19,85 @@ We define the subtype of open sets in a topological space. ## Main Definitions +### Bundled open sets + - `opens α` is the type of open subsets of a topological space `α`. +- `opens.is_basis` is a predicate saying that a set of `opens`s form a topological basis. +- `opens.comap`: preimage of an open set under a continuous map as a `frame_hom`. +- `homeomorph.opens_congr`: order-preserving equivalence between open sets in the domain and the + codomain of a homeomorphism. + +### Bundled open neighborhoods + - `open_nhds_of x` is the type of open subsets of a topological space `α` containing `x : α`. +- `open_nhds_of.comap f x U` is the preimage of open neighborhood `U` of `f x` under `f : C(α, β)`. + +## Main results + +We define order structures on both `opens α` (`complete_structure`, `frame`) and `open_nhds_of x` +(`order_top`, `distrib_lattice`). -/ open filter function order set +open_locale topology variables {ι α β γ : Type*} [topological_space α] [topological_space β] [topological_space γ] namespace topological_space + variable (α) + /-- The type of open subsets of a topological space. -/ -def opens := {s : set α // is_open s} +structure opens := +(carrier : set α) +(is_open' : is_open carrier) variable {α} + namespace opens -instance : has_coe (opens α) (set α) := { coe := subtype.val } -lemma val_eq_coe (U : opens α) : U.1 = ↑U := rfl +instance : set_like (opens α) α := +{ coe := opens.carrier, + coe_injective' := λ ⟨_, _⟩ ⟨_, _⟩ _, by congr; assumption } + +instance : can_lift (set α) (opens α) coe is_open := +⟨λ s h, ⟨⟨s, h⟩, rfl⟩⟩ + +lemma «forall» {p : opens α → Prop} : (∀ U, p U) ↔ ∀ (U : set α) (hU : is_open U), p ⟨U, hU⟩ := +⟨λ h _ _, h _, λ h ⟨U, hU⟩, h _ _⟩ + +@[simp] lemma carrier_eq_coe (U : opens α) : U.1 = ↑U := rfl /-- the coercion `opens α → set α` applied to a pair is the same as taking the first component -/ -lemma coe_mk {α : Type*} [topological_space α] {U : set α} {hU : is_open U} : - ↑(⟨U, hU⟩ : opens α) = U := rfl +@[simp] lemma coe_mk {U : set α} {hU : is_open U} : ↑(⟨U, hU⟩ : opens α) = U := rfl -instance : has_subset (opens α) := -{ subset := λ U V, (U : set α) ⊆ V } +@[simp] lemma mem_mk {x : α} {U : set α} {h : is_open U} : + @has_mem.mem _ (opens α) _ x ⟨U, h⟩ ↔ x ∈ U := iff.rfl -instance : has_mem α (opens α) := -{ mem := λ a U, a ∈ (U : set α) } +-- todo: make it `simp` for a `set_like`? +@[simp] protected lemma nonempty_coe_sort {U : opens α} : nonempty U ↔ (U : set α).nonempty := +set.nonempty_coe_sort -@[simp] lemma subset_coe {U V : opens α} : ((U : set α) ⊆ (V : set α)) = (U ⊆ V) := rfl +@[ext] lemma ext {U V : opens α} (h : (U : set α) = V) : U = V := set_like.coe_injective h +@[simp] lemma coe_inj {U V : opens α} : (U : set α) = V ↔ U = V := set_like.ext'_iff.symm -@[simp] lemma mem_coe {x : α} {U : opens α} : (x ∈ (U : set α)) = (x ∈ U) := rfl +protected lemma is_open (U : opens α) : is_open (U : set α) := U.is_open' -@[simp] lemma mem_mk {x : α} {U : set α} {h : is_open U} : - @has_mem.mem _ _ opens.has_mem x ⟨U, h⟩ ↔ x ∈ U := iff.rfl +@[simp] lemma mk_coe (U : opens α) : mk ↑U U.is_open = U := by { cases U, refl } -@[ext] lemma ext {U V : opens α} (h : (U : set α) = V) : U = V := subtype.ext h -@[ext] lemma ext_iff {U V : opens α} : (U : set α) = V ↔ U = V := subtype.ext_iff.symm +/-- See Note [custom simps projection]. -/ +def simps.coe (U : opens α) : set α := U -instance : partial_order (opens α) := subtype.partial_order _ +initialize_simps_projections opens (carrier → coe) /-- The interior of a set, as an element of `opens`. -/ def interior (s : set α) : opens α := ⟨interior s, is_open_interior⟩ lemma gc : galois_connection (coe : opens α → set α) interior := -λ U s, ⟨λ h, interior_maximal h U.property, λ h, le_trans h interior_subset⟩ - -open order_dual (of_dual to_dual) +λ U s, ⟨λ h, interior_maximal h U.is_open, λ h, le_trans h interior_subset⟩ /-- The galois coinsertion between sets and opens. -/ -def gi : galois_coinsertion subtype.val (@interior α _) := +def gi : galois_coinsertion coe (@interior α _) := { choice := λ s hs, ⟨s, interior_eq_iff_is_open.mp $ le_antisymm interior_subset hs⟩, gc := gc, u_l_le := λ _, interior_subset, @@ -76,17 +105,14 @@ def gi : galois_coinsertion subtype.val (@interior α _) := instance : complete_lattice (opens α) := complete_lattice.copy (galois_coinsertion.lift_complete_lattice gi) -/- le -/ (λ U V, U ⊆ V) rfl +/- le -/ (λ U V, (U : set α) ⊆ V) rfl /- top -/ ⟨univ, is_open_univ⟩ (ext interior_univ.symm) /- bot -/ ⟨∅, is_open_empty⟩ rfl /- sup -/ (λ U V, ⟨↑U ∪ ↑V, U.2.union V.2⟩) rfl -/- inf -/ (λ U V, ⟨↑U ∩ ↑V, U.2.inter V.2⟩) - (funext $ λ U, funext $ λ V, ext (U.2.inter V.2).interior_eq.symm) +/- inf -/ (λ U V, ⟨↑U ∩ ↑V, U.2.inter V.2⟩) (funext₂ $ λ U V, ext (U.2.inter V.2).interior_eq.symm) /- Sup -/ (λ S, ⟨⋃ s ∈ S, ↑s, is_open_bUnion $ λ s _, s.2⟩) (funext $ λ S, ext Sup_image.symm) /- Inf -/ _ rfl -lemma le_def {U V : opens α} : U ≤ V ↔ (U : set α) ≤ (V : set α) := iff.rfl - @[simp] lemma mk_inf_mk {U V : set α} {hU : is_open U} {hV : is_open V} : (⟨U, hU⟩ ⊓ ⟨V, hV⟩ : opens α) = ⟨U ⊓ V, is_open.inter hU hV⟩ := rfl @[simp, norm_cast] lemma coe_inf (s t : opens α) : (↑(s ⊓ t) : set α) = s ∩ t := rfl @@ -103,14 +129,7 @@ map_finset_sup (⟨⟨coe, coe_sup⟩, coe_bot⟩ : sup_bot_hom (opens α) (set (↑(s.inf f) : set α) = s.inf (coe ∘ f) := map_finset_inf (⟨⟨coe, coe_inf⟩, coe_top⟩ : inf_top_hom (opens α) (set α)) _ _ -instance : has_inter (opens α) := ⟨λ U V, U ⊓ V⟩ -instance : has_union (opens α) := ⟨λ U V, U ⊔ V⟩ -instance : has_emptyc (opens α) := ⟨⊥⟩ -instance : inhabited (opens α) := ⟨∅⟩ - -@[simp] lemma inter_eq (U V : opens α) : U ∩ V = U ⊓ V := rfl -@[simp] lemma union_eq (U V : opens α) : U ∪ V = U ⊔ V := rfl -@[simp] lemma empty_eq : (∅ : opens α) = ⊥ := rfl +instance : inhabited (opens α) := ⟨⊥⟩ lemma supr_def {ι} (s : ι → opens α) : (⨆ i, s i) = ⟨⋃ i, s i, is_open_Union $ λ i, (s i).2⟩ := by { ext, simp only [supr, coe_Sup, bUnion_range], refl } @@ -124,7 +143,7 @@ by { rw supr_def, simp } by simp [supr_def] @[simp] theorem mem_supr {ι} {x : α} {s : ι → opens α} : x ∈ supr s ↔ ∃ i, x ∈ s i := -by { rw [←mem_coe], simp, } +by { rw [← set_like.mem_coe], simp, } @[simp] lemma mem_Sup {Us : set (opens α)} {x : α} : x ∈ Sup Us ↔ ∃ u ∈ Us, x ∈ u := by simp_rw [Sup_eq_supr, mem_supr] @@ -142,24 +161,28 @@ lemma open_embedding_of_le {U V : opens α} (i : U ≤ V) : open_range := begin rw set.range_inclusion i, - exact U.property.preimage continuous_subtype_val + exact U.is_open.preimage continuous_subtype_val end, } lemma not_nonempty_iff_eq_bot (U : opens α) : ¬ set.nonempty (U : set α) ↔ U = ⊥ := -by rw [← subtype.coe_injective.eq_iff, opens.coe_bot, ← set.not_nonempty_iff_eq_empty] +by rw [← coe_inj, opens.coe_bot, ← set.not_nonempty_iff_eq_empty] lemma ne_bot_iff_nonempty (U : opens α) : U ≠ ⊥ ↔ set.nonempty (U : set α) := by rw [ne.def, ← opens.not_nonempty_iff_eq_bot, not_not] /-- An open set in the indiscrete topology is either empty or the whole space. -/ lemma eq_bot_or_top {α} [t : topological_space α] (h : t = ⊤) (U : opens α) : U = ⊥ ∨ U = ⊤ := -by { simp_rw ← ext_iff, unfreezingI { subst h }, exact (is_open_top_iff U.1).1 U.2 } +begin + simp only [← coe_inj], + unfreezingI { subst h }, letI : topological_space α := ⊤, + exact (is_open_top_iff _).1 U.2 +end /-- A set of `opens α` is a basis if the set of corresponding sets is a topological basis. -/ def is_basis (B : set (opens α)) : Prop := is_topological_basis ((coe : _ → set α) '' B) lemma is_basis_iff_nbhd {B : set (opens α)} : - is_basis B ↔ ∀ {U : opens α} {x}, x ∈ U → ∃ U' ∈ B, x ∈ U' ∧ U' ⊆ U := + is_basis B ↔ ∀ {U : opens α} {x}, x ∈ U → ∃ U' ∈ B, x ∈ U' ∧ U' ≤ U := begin split; intro h, { rintros ⟨sU, hU⟩ x hx, @@ -168,7 +191,7 @@ begin refine ⟨V, H₁, _⟩, cases V, dsimp at H₂, subst H₂, exact hsV }, { refine is_topological_basis_of_open_of_nhds _ _, - { rintros sU ⟨U, ⟨H₁, rfl⟩⟩, exact U.property }, + { rintros sU ⟨U, ⟨H₁, rfl⟩⟩, exact U.2 }, { intros x sU hx hsU, rcases @h (⟨sU, hsU⟩ : opens α) x hx with ⟨V, hV, H⟩, exact ⟨V, ⟨V, hV, rfl⟩, H⟩ } } @@ -179,9 +202,9 @@ lemma is_basis_iff_cover {B : set (opens α)} : begin split, { intros hB U, - refine ⟨{V : opens α | V ∈ B ∧ V ⊆ U}, λ U hU, hU.left, _⟩, + refine ⟨{V : opens α | V ∈ B ∧ V ≤ U}, λ U hU, hU.left, _⟩, apply ext, - rw [coe_Sup, hB.open_eq_sUnion' U.prop], + rw [coe_Sup, hB.open_eq_sUnion' U.is_open], simp_rw [sUnion_eq_bUnion, Union, supr_and, supr_image], refl }, { intro h, @@ -194,8 +217,8 @@ end /-- If `α` has a basis consisting of compact opens, then an open set in `α` is compact open iff it is a finite union of some elements in the basis -/ -lemma is_compact_open_iff_eq_finite_Union_of_is_basis - {ι : Type*} (b : ι → opens α) (hb : opens.is_basis (set.range b)) +lemma is_basis.is_compact_open_iff_eq_finite_Union + {ι : Type*} (b : ι → opens α) (hb : is_basis (set.range b)) (hb' : ∀ i, is_compact (b i : set α)) (U : set α) : is_compact U ∧ is_open U ↔ ∃ (s : set ι), s.finite ∧ U = ⋃ i ∈ s, b i := begin @@ -215,7 +238,7 @@ begin refine ⟨t, set.subset.trans ht _⟩, rw [coe_finset_sup, finset.sup_eq_supr], refl }, - { obtain ⟨t, ht⟩ := H (λ i, U i) (λ i, (U i).prop) + { obtain ⟨t, ht⟩ := H (λ i, U i) (λ i, (U i).is_open) (by simpa using (show (s : set α) ⊆ ↑(supr U), from hU)), refine ⟨t, set.subset.trans ht _⟩, simp only [set.Union_subset_iff], @@ -225,8 +248,7 @@ end /-- The preimage of an open set, as an open set. -/ def comap (f : C(α, β)) : frame_hom (opens β) (opens α) := { to_fun := λ s, ⟨f ⁻¹' s, s.2.preimage f.continuous⟩, - map_Sup' := λ s, ext $ by simp only [coe_Sup, preimage_Union, coe_mk, mem_image, Union_exists, - bUnion_and', Union_Union_eq_right], + map_Sup' := λ s, ext $ by simp only [coe_Sup, preimage_Union, bUnion_image, coe_mk], map_inf' := λ a b, rfl, map_top' := rfl } @@ -238,8 +260,6 @@ order_hom_class.mono (comap f) h @[simp] lemma coe_comap (f : C(α, β)) (U : opens β) : ↑(comap f U) = f ⁻¹' U := rfl -@[simp] lemma comap_val (f : C(α, β)) (U : opens β) : (comap f U).1 = f ⁻¹' U := rfl - protected lemma comap_comp (g : C(β, γ)) (f : C(α, β)) : comap (g.comp f) = (comap f).comp (comap g) := rfl @@ -249,28 +269,88 @@ protected lemma comap_comap (g : C(β, γ)) (f : C(α, β)) (U : opens γ) : lemma comap_injective [t0_space β] : injective (comap : C(α, β) → frame_hom (opens β) (opens α)) := λ f g h, continuous_map.ext $ λ a, inseparable.eq $ inseparable_iff_forall_open.2 $ λ s hs, have comap f ⟨s, hs⟩ = comap g ⟨s, hs⟩, from fun_like.congr_fun h ⟨_, hs⟩, -show a ∈ f ⁻¹' s ↔ a ∈ g ⁻¹' s, from set.ext_iff.1 (ext_iff.2 this) a +show a ∈ f ⁻¹' s ↔ a ∈ g ⁻¹' s, from set.ext_iff.1 (coe_inj.2 this) a -/-- A homeomorphism induces an equivalence on open sets, by taking comaps. -/ -@[simp] protected def equiv (f : α ≃ₜ β) : opens α ≃ opens β := +/-- A homeomorphism induces an order-preserving equivalence on open sets, by taking comaps. -/ +@[simps apply { fully_applied := ff }] +def _root_.homeomorph.opens_congr (f : α ≃ₜ β) : opens α ≃o opens β := { to_fun := opens.comap f.symm.to_continuous_map, inv_fun := opens.comap f.to_continuous_map, left_inv := by { intro U, ext1, exact f.to_equiv.preimage_symm_preimage _ }, - right_inv := by { intro U, ext1, exact f.to_equiv.symm_preimage_preimage _ } } + right_inv := by { intro U, ext1, exact f.to_equiv.symm_preimage_preimage _ }, + map_rel_iff' := λ U V, by simp only [← set_like.coe_subset_coe]; + exact f.symm.surjective.preimage_subset_preimage_iff } -/-- A homeomorphism induces an order isomorphism on open sets, by taking comaps. -/ -@[simp] protected def order_iso (f : α ≃ₜ β) : opens α ≃o opens β := -{ to_equiv := opens.equiv f, - map_rel_iff' := λ U V, f.symm.surjective.preimage_subset_preimage_iff } +@[simp] lemma _root_.homeomorph.opens_congr_symm (f : α ≃ₜ β) : + f.opens_congr.symm = f.symm.opens_congr := +rfl + +instance [finite α] : finite (opens α) := finite.of_injective _ set_like.coe_injective end opens /-- The open neighborhoods of a point. See also `opens` or `nhds`. -/ -def open_nhds_of (x : α) : Type* := { s : set α // is_open s ∧ x ∈ s } +structure open_nhds_of (x : α) extends opens α := +(mem' : x ∈ carrier) + +namespace open_nhds_of + +variables {x : α} + +lemma to_opens_injective : injective (to_opens : open_nhds_of x → opens α) +| ⟨_, _⟩ ⟨_, _⟩ rfl := rfl + +instance : set_like (open_nhds_of x) α := +{ coe := λ U, U.1, + coe_injective' := set_like.coe_injective.comp to_opens_injective } + +instance can_lift_set : can_lift (set α) (open_nhds_of x) coe (λ s, is_open s ∧ x ∈ s) := +⟨λ s hs, ⟨⟨⟨s, hs.1⟩, hs.2⟩, rfl⟩⟩ + +protected lemma mem (U : open_nhds_of x) : x ∈ U := U.mem' +protected lemma is_open (U : open_nhds_of x) : is_open (U : set α) := U.is_open' + +instance : order_top (open_nhds_of x) := +{ top := ⟨⊤, set.mem_univ _⟩, + le_top := λ _, subset_univ _ } + +instance : inhabited (open_nhds_of x) := ⟨⊤⟩ + +instance : has_inf (open_nhds_of x) := ⟨λ U V, ⟨U.1 ⊓ V.1, U.2, V.2⟩⟩ -instance open_nhds_of.inhabited {α : Type*} [topological_space α] (x : α) : - inhabited (open_nhds_of x) := ⟨⟨set.univ, is_open_univ, set.mem_univ _⟩⟩ +instance : has_sup (open_nhds_of x) := ⟨λ U V, ⟨U.1 ⊔ V.1, or.inl U.2⟩⟩ -instance [finite α] : finite (opens α) := subtype.finite +instance : distrib_lattice (open_nhds_of x) := +to_opens_injective.distrib_lattice _ (λ _ _, rfl) (λ _ _, rfl) + +lemma basis_nhds : (𝓝 x).has_basis (λ U : open_nhds_of x, true) coe := +(nhds_basis_opens x).to_has_basis (λ U hU, ⟨⟨⟨U, hU.2⟩, hU.1⟩, trivial, subset.rfl⟩) + (λ U _, ⟨U, ⟨⟨U.mem, U.is_open⟩, subset.rfl⟩⟩) + +/-- Preimage of an open neighborhood of `f x` under a continuous map `f` as a `lattice_hom`. -/ +def comap (f : C(α, β)) (x : α) : lattice_hom (open_nhds_of (f x)) (open_nhds_of x) := +{ to_fun := λ U, ⟨opens.comap f U.1, U.mem⟩, + map_sup' := λ U V, rfl, + map_inf' := λ U V, rfl } + +end open_nhds_of end topological_space + +namespace tactic + +namespace auto_cases + +/-- Find an `auto_cases_tac` which matches `topological_space.opens`. -/ +meta def opens_find_tac : expr → option auto_cases_tac +| `(topological_space.opens _) := tac_cases +| _ := none + +end auto_cases + +/-- A version of `tactic.auto_cases` that works for `topological_space.opens`. -/ +@[hint_tactic] +meta def auto_cases_opens : tactic string := +auto_cases tactic.auto_cases.opens_find_tac + +end tactic diff --git a/src/topology/sheaves/presheaf.lean b/src/topology/sheaves/presheaf.lean index 1a39ca3fc7a89..70726dfecf1cc 100644 --- a/src/topology/sheaves/presheaf.lean +++ b/src/topology/sheaves/presheaf.lean @@ -176,7 +176,14 @@ by { dsimp [id], simp, } local attribute [tidy] tactic.op_induction' @[simp, priority 990] lemma id_hom_app (U) : - (id ℱ).hom.app U = ℱ.map (eq_to_hom (opens.op_map_id_obj U)) := by tidy + (id ℱ).hom.app U = ℱ.map (eq_to_hom (opens.op_map_id_obj U)) := +begin + -- was `tidy` + induction U using opposite.rec, + cases U, + rw [id_hom_app'], + congr +end @[simp] lemma id_inv_app' (U) (p) : (id ℱ).inv.app (op ⟨U, p⟩) = ℱ.map (𝟙 (op ⟨U, p⟩)) := by { dsimp [id], simp, } diff --git a/src/topology/sheaves/punit.lean b/src/topology/sheaves/punit.lean index c6962e31579cc..c036317c91fdf 100644 --- a/src/topology/sheaves/punit.lean +++ b/src/topology/sheaves/punit.lean @@ -29,7 +29,7 @@ lemma is_sheaf_of_is_terminal_of_indiscrete {X : Top.{w}} (hind : X.str = ⊤) ( { convert presieve.is_sheaf_for_top_sieve _, rw ←sieve.id_mem_iff_eq_top, have := (U.eq_bot_or_top hind).resolve_left hne, subst this, obtain he | ⟨⟨x⟩⟩ := is_empty_or_nonempty X, - { exact (hne $ topological_space.opens.ext_iff.1 $ set.univ_eq_empty_iff.2 he).elim }, + { exact (hne $ set_like.ext'_iff.2 $ set.univ_eq_empty_iff.2 he).elim }, obtain ⟨U, f, hf, hm⟩ := hs x trivial, obtain rfl | rfl := U.eq_bot_or_top hind, { cases hm }, { convert hf } }, diff --git a/src/topology/sheaves/sheaf_condition/sites.lean b/src/topology/sheaves/sheaf_condition/sites.lean index ebc32d829f9d0..815f909b02492 100644 --- a/src/topology/sheaves/sheaf_condition/sites.lean +++ b/src/topology/sheaves/sheaf_condition/sites.lean @@ -59,7 +59,7 @@ begin intro f, exact f.2.1.le, }, intros x hxU, - rw [opens.mem_coe, opens.mem_supr], + rw [opens.mem_supr], obtain ⟨V, iVU, ⟨W, iVW, iWU, hiWU, -⟩, hxV⟩ := hR x hxU, exact ⟨⟨W, ⟨iWU, hiWU⟩⟩, iVW.le hxV⟩, end @@ -183,11 +183,11 @@ variables {X : Top.{w}} {ι : Type*} {B : ι → opens X} variables (F : X.presheaf C) (F' : sheaf C X) (h : opens.is_basis (set.range B)) /-- The empty component of a sheaf is terminal -/ -def is_terminal_of_empty (F : sheaf C X) : limits.is_terminal (F.val.obj (op ∅)) := -F.is_terminal_of_bot_cover ∅ (by tidy) +def is_terminal_of_empty (F : sheaf C X) : limits.is_terminal (F.val.obj (op ⊥)) := +F.is_terminal_of_bot_cover ⊥ (by tidy) /-- A variant of `is_terminal_of_empty` that is easier to `apply`. -/ -def is_terminal_of_eq_empty (F : X.sheaf C) {U : opens X} (h : U = ∅) : +def is_terminal_of_eq_empty (F : X.sheaf C) {U : opens X} (h : U = ⊥) : limits.is_terminal (F.val.obj (op U)) := by convert F.is_terminal_of_empty diff --git a/src/topology/sheaves/stalks.lean b/src/topology/sheaves/stalks.lean index 772a1e0c892c7..7bf0501687132 100644 --- a/src/topology/sheaves/stalks.lean +++ b/src/topology/sheaves/stalks.lean @@ -193,7 +193,7 @@ lemma stalk_pushforward_iso_of_open_embedding {f : X ⟶ Y} (hf : open_embedding { intro U, refine F.map_iso (eq_to_iso _), dsimp only [functor.op], - exact congr_arg op (subtype.eq $ set.preimage_image_eq (unop U).1.1 hf.inj) }, + exact congr_arg op (opens.ext $ set.preimage_image_eq (unop U).1.1 hf.inj) }, { intros U V i, erw [← F.map_comp, ← F.map_comp], congr } }, { ext U, rw ← iso.comp_inv_eq, @@ -299,12 +299,12 @@ end @[simp, reassoc, elementwise] lemma germ_stalk_specializes (F : X.presheaf C) {U : opens X} {y : U} {x : X} (h : x ⤳ y) : F.germ y ≫ F.stalk_specializes h = - F.germ ⟨x, specializes_iff_forall_open.mp h _ U.2 y.prop⟩ := colimit.ι_desc _ _ + F.germ (⟨x, h.mem_open U.is_open y.prop⟩ : U) := colimit.ι_desc _ _ @[simp, reassoc, elementwise] lemma germ_stalk_specializes' (F : X.presheaf C) {U : opens X} {x y : X} (h : x ⤳ y) (hy : y ∈ U) : F.germ ⟨y, hy⟩ ≫ F.stalk_specializes h = - F.germ ⟨x, specializes_iff_forall_open.mp h _ U.2 hy⟩ := colimit.ι_desc _ _ + F.germ ⟨x, h.mem_open U.is_open hy⟩ := colimit.ι_desc _ _ @[simp] lemma stalk_specializes_refl {C : Type*} [category C] [limits.has_colimits C] @@ -417,7 +417,7 @@ begin -- neighborhoods form a cover of `U`. apply F.eq_of_locally_eq' V U i₁, { intros x hxU, - rw [opens.mem_coe, opens.mem_supr], + rw [opens.mem_supr], exact ⟨⟨x, hxU⟩, m ⟨x, hxU⟩⟩ }, { intro x, rw [heq, subsingleton.elim (i₁ x) (i₂ x)] } @@ -481,7 +481,7 @@ begin -- These neighborhoods clearly cover all of `U`. have V_cover : U ≤ supr V, { intros x hxU, - rw [opens.mem_coe, opens.mem_supr], + rw [opens.mem_supr], exact ⟨⟨x, hxU⟩, mV ⟨x, hxU⟩⟩ }, -- Since `F` is a sheaf, we can glue all the local preimages together to get a global preimage. obtain ⟨s, s_spec, -⟩ := F.exists_unique_gluing' V U iVU V_cover sf _, From 8c8c544bf24ced19b1e76c34bb3262bdae620f82 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 9 Feb 2023 19:18:05 +0000 Subject: [PATCH 0455/1291] chore(topology/algebra): make the divisor argument of `div_const` explicit (#18411) This is using the rule "parameters which do not appear in the types of other parameters should be explicit". In particular, this means that it is easier to use proofs of the form `convert hf.div_const c`, without having to omit `c` and hope that Lean can guess it. This makes the argument explicit for: * `cont_diff_within_at.div_const` * `cont_diff_at.div_const` * `cont_diff_on.div_const` * `cont_diff.div_const` * `filter.tendsto.div_const` * `continuous_at.div_const` * `continuous_within_at.div_const` * `continuous_on.div_const` * `continuous.div_const` * `differentiable_within_at.div_const` * `differentiable_at.div_const` * `differentiable_on.div_const` * `differentiable.div_const` * `deriv_within_div_const` It was already explicit for: * `filter.tendsto.div_const'` * `has_sum.div_const` * `summable.div_const` * `integrable.div_const` * `has_deriv_at.div_const` * `has_deriv_within_at.div_const` * `has_strict_deriv_at.div_const` * `periodic.div_const` * `measurable.div_const` * `ae_measurable.div_const` * The `mul` variants of many of the above --- src/analysis/calculus/bump_function_findim.lean | 4 ++-- src/analysis/calculus/bump_function_inner.lean | 4 ++-- src/analysis/calculus/cont_diff.lean | 14 +++++++------- src/analysis/calculus/deriv.lean | 14 +++++++------- src/analysis/normed/field/basic.lean | 2 +- src/analysis/special_functions/integrals.lean | 4 ++-- src/analysis/special_functions/log/deriv.lean | 2 +- .../special_functions/trigonometric/deriv.lean | 8 ++++---- src/topology/algebra/group/basic.lean | 4 ++-- src/topology/algebra/group_with_zero.lean | 12 ++++++------ 10 files changed, 34 insertions(+), 34 deletions(-) diff --git a/src/analysis/calculus/bump_function_findim.lean b/src/analysis/calculus/bump_function_findim.lean index 848daeca1936b..66e5922f13c04 100644 --- a/src/analysis/calculus/bump_function_findim.lean +++ b/src/analysis/calculus/bump_function_findim.lean @@ -217,7 +217,7 @@ begin from A.exists_smooth_support_eq, have B : ∀ x, f x ∈ Icc (0 : ℝ) 1 := λ x, f_range (mem_range_self x), refine ⟨λ x, (f x + f (-x)) / 2, _, _, _, _⟩, - { exact (f_smooth.add (f_smooth.comp cont_diff_neg)).div_const }, + { exact (f_smooth.add (f_smooth.comp cont_diff_neg)).div_const _ }, { assume x, split, { linarith [(B x).1, (B (-x)).1] }, @@ -513,7 +513,7 @@ begin dsimp only, linarith, }, { apply cont_diff_on.smul _ cont_diff_on_snd, - refine (cont_diff_on_fst.add cont_diff_on_const).div_const.inv _, + refine ((cont_diff_on_fst.add cont_diff_on_const).div_const _).inv _, rintros ⟨R, x⟩ ⟨(hR : 1 < R), hx⟩, apply ne_of_gt, dsimp only, diff --git a/src/analysis/calculus/bump_function_inner.lean b/src/analysis/calculus/bump_function_inner.lean index 8a8987b7af8cd..47c1955e56285 100644 --- a/src/analysis/calculus/bump_function_inner.lean +++ b/src/analysis/calculus/bump_function_inner.lean @@ -505,10 +505,10 @@ lemma nonneg_normed (x : E) : 0 ≤ f.normed μ x := div_nonneg f.nonneg $ integral_nonneg f.nonneg' lemma cont_diff_normed {n : ℕ∞} : cont_diff ℝ n (f.normed μ) := -f.cont_diff.div_const +f.cont_diff.div_const _ lemma continuous_normed : continuous (f.normed μ) := -f.continuous.div_const +f.continuous.div_const _ lemma normed_sub (x : E) : f.normed μ (c - x) = f.normed μ (c + x) := by simp_rw [f.normed_def, f.sub] diff --git a/src/analysis/calculus/cont_diff.lean b/src/analysis/calculus/cont_diff.lean index 3f8f95d40cdf6..5274779ceabab 100644 --- a/src/analysis/calculus/cont_diff.lean +++ b/src/analysis/calculus/cont_diff.lean @@ -2861,20 +2861,20 @@ lemma cont_diff_on.pow {f : E → 𝔸} (hf : cont_diff_on 𝕜 n f s) (m : ℕ) cont_diff_on 𝕜 n (λ y, f y ^ m) s := λ y hy, (hf y hy).pow m -lemma cont_diff_within_at.div_const {f : E → 𝕜'} {n} {c : 𝕜'} - (hf : cont_diff_within_at 𝕜 n f s x) : +lemma cont_diff_within_at.div_const {f : E → 𝕜'} {n} + (hf : cont_diff_within_at 𝕜 n f s x) (c : 𝕜') : cont_diff_within_at 𝕜 n (λ x, f x / c) s x := by simpa only [div_eq_mul_inv] using hf.mul cont_diff_within_at_const -lemma cont_diff_at.div_const {f : E → 𝕜'} {n} {c : 𝕜'} (hf : cont_diff_at 𝕜 n f x) : +lemma cont_diff_at.div_const {f : E → 𝕜'} {n} (hf : cont_diff_at 𝕜 n f x) (c : 𝕜') : cont_diff_at 𝕜 n (λ x, f x / c) x := -hf.div_const +hf.div_const c -lemma cont_diff_on.div_const {f : E → 𝕜'} {n} {c : 𝕜'} (hf : cont_diff_on 𝕜 n f s) : +lemma cont_diff_on.div_const {f : E → 𝕜'} {n} (hf : cont_diff_on 𝕜 n f s) (c : 𝕜') : cont_diff_on 𝕜 n (λ x, f x / c) s := -λ x hx, (hf x hx).div_const +λ x hx, (hf x hx).div_const c -lemma cont_diff.div_const {f : E → 𝕜'} {n} {c : 𝕜'} (hf : cont_diff 𝕜 n f) : +lemma cont_diff.div_const {f : E → 𝕜'} {n} (hf : cont_diff 𝕜 n f) (c : 𝕜') : cont_diff 𝕜 n (λ x, f x / c) := by simpa only [div_eq_mul_inv] using hf.mul cont_diff_const diff --git a/src/analysis/calculus/deriv.lean b/src/analysis/calculus/deriv.lean index 1665fcc1ecb7c..d3f5f5e96cce2 100644 --- a/src/analysis/calculus/deriv.lean +++ b/src/analysis/calculus/deriv.lean @@ -1639,23 +1639,23 @@ lemma has_strict_deriv_at.div_const (hc : has_strict_deriv_at c c' x) (d : 𝕜' has_strict_deriv_at (λ x, c x / d) (c' / d) x := by simpa only [div_eq_mul_inv] using hc.mul_const d⁻¹ -lemma differentiable_within_at.div_const (hc : differentiable_within_at 𝕜 c s x) {d : 𝕜'} : +lemma differentiable_within_at.div_const (hc : differentiable_within_at 𝕜 c s x) (d : 𝕜') : differentiable_within_at 𝕜 (λx, c x / d) s x := (hc.has_deriv_within_at.div_const _).differentiable_within_at -@[simp] lemma differentiable_at.div_const (hc : differentiable_at 𝕜 c x) {d : 𝕜'} : +@[simp] lemma differentiable_at.div_const (hc : differentiable_at 𝕜 c x) (d : 𝕜') : differentiable_at 𝕜 (λ x, c x / d) x := (hc.has_deriv_at.div_const _).differentiable_at -lemma differentiable_on.div_const (hc : differentiable_on 𝕜 c s) {d : 𝕜'} : +lemma differentiable_on.div_const (hc : differentiable_on 𝕜 c s) (d : 𝕜') : differentiable_on 𝕜 (λx, c x / d) s := -λ x hx, (hc x hx).div_const +λ x hx, (hc x hx).div_const d -@[simp] lemma differentiable.div_const (hc : differentiable 𝕜 c) {d : 𝕜'} : +@[simp] lemma differentiable.div_const (hc : differentiable 𝕜 c) (d : 𝕜') : differentiable 𝕜 (λx, c x / d) := -λ x, (hc x).div_const +λ x, (hc x).div_const d -lemma deriv_within_div_const (hc : differentiable_within_at 𝕜 c s x) {d : 𝕜'} +lemma deriv_within_div_const (hc : differentiable_within_at 𝕜 c s x) (d : 𝕜') (hxs : unique_diff_within_at 𝕜 s x) : deriv_within (λx, c x / d) s x = (deriv_within c s x) / d := by simp [div_eq_inv_mul, deriv_within_const_mul, hc, hxs] diff --git a/src/analysis/normed/field/basic.lean b/src/analysis/normed/field/basic.lean index 6b16b153c1355..efa11872c9429 100644 --- a/src/analysis/normed/field/basic.lean +++ b/src/analysis/normed/field/basic.lean @@ -483,7 +483,7 @@ begin ... ≤ ‖r - e‖ / ‖r‖ / ε : div_le_div_of_le_left (div_nonneg (norm_nonneg _) (norm_nonneg _)) ε0 he.le }, refine squeeze_zero' (eventually_of_forall $ λ _, norm_nonneg _) this _, - refine (continuous_const.sub continuous_id).norm.div_const.div_const.tendsto' _ _ _, + refine (((continuous_const.sub continuous_id).norm.div_const _).div_const _).tendsto' _ _ _, simp, end diff --git a/src/analysis/special_functions/integrals.lean b/src/analysis/special_functions/integrals.lean index c657a727b5564..6bcc48d15b912 100644 --- a/src/analysis/special_functions/integrals.lean +++ b/src/analysis/special_functions/integrals.lean @@ -69,7 +69,7 @@ begin field_simp [(by linarith : r + 1 ≠ 0)], ring, }, apply integrable_on_deriv_of_nonneg hc _ hderiv, { intros x hx, apply rpow_nonneg_of_nonneg hx.1.le, }, - { refine (continuous_on_id.rpow_const _).div_const, intros x hx, right, linarith } }, + { refine (continuous_on_id.rpow_const _).div_const _, intros x hx, right, linarith } }, intro c, rcases le_total 0 c with hc|hc, { exact this c hc }, { rw [interval_integrable.iff_comp_neg, neg_zero], @@ -296,7 +296,7 @@ begin ring }, intro c, apply integral_eq_sub_of_has_deriv_right, - { refine (complex.continuous_of_real_cpow_const _).div_const.continuous_on, + { refine ((complex.continuous_of_real_cpow_const _).div_const _).continuous_on, rwa [complex.add_re, complex.one_re, ←neg_lt_iff_pos_add] }, { refine λ x hx, (has_deriv_at_of_real_cpow _ _).has_deriv_within_at, { rcases le_total c 0 with hc | hc, diff --git a/src/analysis/special_functions/log/deriv.lean b/src/analysis/special_functions/log/deriv.lean index 23629e5f7f7c5..8db6780d53702 100644 --- a/src/analysis/special_functions/log/deriv.lean +++ b/src/analysis/special_functions/log/deriv.lean @@ -248,7 +248,7 @@ begin suffices : tendsto (λ (t : ℕ), |x| ^ (t + 1) / (1 - |x|)) at_top (𝓝 (|x| * 0 / (1 - |x|))), by simpa, simp only [pow_succ], - refine (tendsto_const_nhds.mul _).div_const, + refine (tendsto_const_nhds.mul _).div_const _, exact tendsto_pow_at_top_nhds_0_of_lt_1 (abs_nonneg _) h }, show summable (λ (n : ℕ), x ^ (n + 1) / (n + 1)), { refine summable_of_norm_bounded _ (summable_geometric_of_lt_1 (abs_nonneg _) h) (λ i, _), diff --git a/src/analysis/special_functions/trigonometric/deriv.lean b/src/analysis/special_functions/trigonometric/deriv.lean index e6bf6f493a075..5edde7a39cbab 100644 --- a/src/analysis/special_functions/trigonometric/deriv.lean +++ b/src/analysis/special_functions/trigonometric/deriv.lean @@ -44,7 +44,7 @@ lemma has_deriv_at_sin (x : ℂ) : has_deriv_at sin (cos x) x := lemma cont_diff_sin {n} : cont_diff ℂ n sin := (((cont_diff_neg.mul cont_diff_const).cexp.sub - (cont_diff_id.mul cont_diff_const).cexp).mul cont_diff_const).div_const + (cont_diff_id.mul cont_diff_const).cexp).mul cont_diff_const).div_const _ lemma differentiable_sin : differentiable ℂ sin := λx, (has_deriv_at_sin x).differentiable_at @@ -72,7 +72,7 @@ lemma has_deriv_at_cos (x : ℂ) : has_deriv_at cos (-sin x) x := lemma cont_diff_cos {n} : cont_diff ℂ n cos := ((cont_diff_id.mul cont_diff_const).cexp.add - (cont_diff_neg.mul cont_diff_const).cexp).div_const + (cont_diff_neg.mul cont_diff_const).cexp).div_const _ lemma differentiable_cos : differentiable ℂ cos := λx, (has_deriv_at_cos x).differentiable_at @@ -101,7 +101,7 @@ lemma has_deriv_at_sinh (x : ℂ) : has_deriv_at sinh (cosh x) x := (has_strict_deriv_at_sinh x).has_deriv_at lemma cont_diff_sinh {n} : cont_diff ℂ n sinh := -(cont_diff_exp.sub cont_diff_neg.cexp).div_const +(cont_diff_exp.sub cont_diff_neg.cexp).div_const _ lemma differentiable_sinh : differentiable ℂ sinh := λx, (has_deriv_at_sinh x).differentiable_at @@ -127,7 +127,7 @@ lemma has_deriv_at_cosh (x : ℂ) : has_deriv_at cosh (sinh x) x := (has_strict_deriv_at_cosh x).has_deriv_at lemma cont_diff_cosh {n} : cont_diff ℂ n cosh := -(cont_diff_exp.add cont_diff_neg.cexp).div_const +(cont_diff_exp.add cont_diff_neg.cexp).div_const _ lemma differentiable_cosh : differentiable ℂ cosh := λx, (has_deriv_at_cosh x).differentiable_at diff --git a/src/topology/algebra/group/basic.lean b/src/topology/algebra/group/basic.lean index c2b5c7ed806e5..5b5dd64a855bd 100644 --- a/src/topology/algebra/group/basic.lean +++ b/src/topology/algebra/group/basic.lean @@ -832,8 +832,8 @@ lemma filter.tendsto.const_div' (b : G) {c : G} {f : α → G} {l : filter α} tendsto_const_nhds.div' h @[to_additive sub_const] -lemma filter.tendsto.div_const' (b : G) {c : G} {f : α → G} {l : filter α} - (h : tendsto f l (𝓝 c)) : tendsto (λ k : α, f k / b) l (𝓝 (c / b)) := +lemma filter.tendsto.div_const' {c : G} {f : α → G} {l : filter α} + (h : tendsto f l (𝓝 c)) (b : G) : tendsto (λ k : α, f k / b) l (𝓝 (c / b)) := h.div' tendsto_const_nhds variables [topological_space α] {f g : α → G} {s : set α} {x : α} diff --git a/src/topology/algebra/group_with_zero.lean b/src/topology/algebra/group_with_zero.lean index 93b8fdccf51e8..f8f1f47ac7c9b 100644 --- a/src/topology/algebra/group_with_zero.lean +++ b/src/topology/algebra/group_with_zero.lean @@ -48,24 +48,24 @@ section div_const variables [group_with_zero G₀] [topological_space G₀] [has_continuous_mul G₀] {f : α → G₀} {s : set α} {l : filter α} -lemma filter.tendsto.div_const {x y : G₀} (hf : tendsto f l (𝓝 x)) : +lemma filter.tendsto.div_const {x : G₀} (hf : tendsto f l (𝓝 x)) (y : G₀) : tendsto (λa, f a / y) l (𝓝 (x / y)) := by simpa only [div_eq_mul_inv] using hf.mul tendsto_const_nhds variables [topological_space α] -lemma continuous_at.div_const {a : α} (hf : continuous_at f a) {y : G₀} : +lemma continuous_at.div_const {a : α} (hf : continuous_at f a) (y : G₀) : continuous_at (λ x, f x / y) a := by simpa only [div_eq_mul_inv] using hf.mul continuous_at_const -lemma continuous_within_at.div_const {a} (hf : continuous_within_at f s a) {y : G₀} : +lemma continuous_within_at.div_const {a} (hf : continuous_within_at f s a) (y : G₀) : continuous_within_at (λ x, f x / y) s a := -hf.div_const +hf.div_const _ -lemma continuous_on.div_const (hf : continuous_on f s) {y : G₀} : continuous_on (λ x, f x / y) s := +lemma continuous_on.div_const (hf : continuous_on f s) (y : G₀) : continuous_on (λ x, f x / y) s := by simpa only [div_eq_mul_inv] using hf.mul continuous_on_const -@[continuity] lemma continuous.div_const (hf : continuous f) {y : G₀} : +@[continuity] lemma continuous.div_const (hf : continuous f) (y : G₀) : continuous (λ x, f x / y) := by simpa only [div_eq_mul_inv] using hf.mul continuous_const From 735b22f8f9ff9792cf4212d7cb051c4c994bc685 Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Fri, 10 Feb 2023 10:20:57 +0000 Subject: [PATCH 0456/1291] feat(analysis/fourier): Riemann-Lebesgue lemma (part 1) (#18361) Add the Riemann-Lebesgue lemma for continuous and compactly-supported functions. Co-authored-by: sgouezel --- src/analysis/complex/basic.lean | 4 + .../fourier/riemann_lebesgue_lemma.lean | 145 ++++++++++++++++++ src/analysis/normed/group/basic.lean | 17 ++ src/measure_theory/integral/set_integral.lean | 17 ++ .../continuous_function/zero_at_infty.lean | 2 +- src/topology/uniform_space/compact.lean | 29 +++- 6 files changed, 209 insertions(+), 5 deletions(-) create mode 100644 src/analysis/fourier/riemann_lebesgue_lemma.lean diff --git a/src/analysis/complex/basic.lean b/src/analysis/complex/basic.lean index c83008b39af8b..6d82754acffc0 100644 --- a/src/analysis/complex/basic.lean +++ b/src/analysis/complex/basic.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ import data.complex.module +import data.complex.exponential import data.is_R_or_C.basic import topology.algebra.module.infinite_sum import topology.instances.real_vector_space @@ -43,6 +44,9 @@ instance : has_norm ℂ := ⟨abs⟩ @[simp] lemma norm_eq_abs (z : ℂ) : ‖z‖ = abs z := rfl +lemma norm_exp_of_real_mul_I (t : ℝ) : ‖exp (t * I)‖ = 1 := +by simp only [norm_eq_abs, abs_exp_of_real_mul_I] + instance : normed_add_comm_group ℂ := add_group_norm.to_normed_add_comm_group { map_zero' := map_zero abs, diff --git a/src/analysis/fourier/riemann_lebesgue_lemma.lean b/src/analysis/fourier/riemann_lebesgue_lemma.lean new file mode 100644 index 0000000000000..4046ebd59f3f0 --- /dev/null +++ b/src/analysis/fourier/riemann_lebesgue_lemma.lean @@ -0,0 +1,145 @@ +/- +Copyright (c) 2022 David Loeffler. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: David Loeffler +-/ +import measure_theory.function.lp_space +import measure_theory.function.continuous_map_dense +import measure_theory.integral.interval_integral +import measure_theory.integral.integral_eq_improper +import measure_theory.group.integration +import topology.continuous_function.zero_at_infty + +/-! +# The Riemann-Lebesgue Lemma + +In this file we prove a weak form of the Riemann-Lebesgue lemma, stating that for any +compactly-supported continuous function `f` on `ℝ` (valued in some complete normed space `E`), the +integral + +`∫ (x : ℝ), exp (↑(t * x) * I) • f x` + +tends to zero as `t → ∞`. (The actual lemma is that this holds for all `L¹` functions `f`, which +follows from the result proved here together with the fact that continuous, compactly-supported +functions are dense in `L¹(ℝ)`, which will be proved in a future iteration.) + +## Main results + +- `tendsto_integral_mul_exp_at_top_of_continuous_compact_support`: the Riemann-Lebesgue lemma for + continuous compactly-supported functions on `ℝ`. +-/ + +open measure_theory filter complex set +open_locale filter topology real ennreal + +section continuous_compact_support + +variables {E : Type*} [normed_add_comm_group E] [normed_space ℂ E] {f : ℝ → E} + +/-- The integrand in the Riemann-Lebesgue lemma is integrable. -/ +lemma fourier_integrand_integrable (hf : integrable f) (t : ℝ) : + integrable (λ x:ℝ, exp (↑(t * x) * I) • f x) := +begin + rw ←integrable_norm_iff, + simp_rw [norm_smul, norm_exp_of_real_mul_I, one_mul], + exacts [hf.norm, (continuous.ae_strongly_measurable (by continuity)).smul hf.1], +end + +variable [complete_space E] + +/-- Shifting `f` by `π / t` negates the integral in the Riemann-Lebesgue lemma. -/ +lemma fourier_integral_half_period_translate {t : ℝ} (ht : t ≠ 0) : + ∫ x:ℝ, exp (↑(t * x) * I) • f (x + π / t) = -∫ x:ℝ, exp (↑(t * x) * I) • f x := +begin + have : (λ x:ℝ, exp (↑(t * x) * I) • f (x + π / t)) = + (λ x:ℝ, (λ y:ℝ, -exp (↑(t * y) * I) • f y) (x + π / t)), + { ext1 x, dsimp only, + rw [of_real_mul, of_real_mul, of_real_add, mul_add, add_mul, exp_add, ←neg_mul], + replace ht := complex.of_real_ne_zero.mpr ht, + have : ↑t * ↑(π / t) * I = π * I, by { field_simp, ring }, + rw [this, exp_pi_mul_I], ring_nf, }, + rw [this, integral_add_right_eq_self], + simp_rw [neg_smul, integral_neg], +end + +/-- Rewrite the Riemann-Lebesgue integral in a form that allows us to use uniform continuity. -/ +lemma fourier_integral_eq_half_sub_half_period_translate + {t : ℝ} (ht : t ≠ 0) (hf : integrable f) : + ∫ x:ℝ, exp (↑(t * x) * I) • f x = + (1 / (2 : ℂ)) • ∫ x:ℝ, exp (↑(t * x) * I) • (f x - f (x + π / t)) := +begin + simp_rw [smul_sub], + rw [integral_sub, fourier_integral_half_period_translate ht, sub_eq_add_neg, neg_neg, + ←two_smul ℂ _, ←@smul_assoc _ _ _ _ _ _ (is_scalar_tower.left ℂ), smul_eq_mul], + norm_num, + exacts [fourier_integrand_integrable hf t, + fourier_integrand_integrable (hf.comp_add_right (π / t)) t], +end + +/-- Riemann-Lebesgue Lemma for continuous and compactly-supported functions: the integral +`∫ x, exp (t * x * I) • f x` tends to 0 as `t` gets large. -/ +lemma tendsto_integral_mul_exp_at_top_of_continuous_compact_support + (hf1 : continuous f) (hf2 : has_compact_support f) : + tendsto (λ t:ℝ, ∫ x:ℝ, exp (↑(t * x) * I) • f x) at_top (𝓝 0) := +begin + simp_rw [normed_add_comm_group.tendsto_nhds_zero, eventually_at_top, ge_iff_le], + intros ε hε, + -- Extract an explicit candidate bound on `t` from uniform continuity. + obtain ⟨R, hR1, hR2⟩ := hf2.exists_pos_le_norm, + obtain ⟨δ, hδ1, hδ2⟩ := metric.uniform_continuous_iff.mp + (hf2.uniform_continuous_of_continuous hf1) (ε / (1 + 2 * R)) (div_pos hε (by positivity)), + refine ⟨max π (1 + π / δ), λ t ht, _⟩, + have tpos : 0 < t := lt_of_lt_of_le real.pi_pos ((le_max_left _ _).trans ht), + -- Rewrite integral in terms of `f x - f (x + π / t)`. + rw fourier_integral_eq_half_sub_half_period_translate (lt_of_lt_of_le + (lt_max_of_lt_left real.pi_pos) ht).ne' (hf1.integrable_of_has_compact_support hf2), + rw [norm_smul, norm_eq_abs, ←complex.of_real_one, ←of_real_bit0, ←of_real_div, + complex.abs_of_nonneg one_half_pos.le], + have : ε = (1 / 2) * (2 * ε), by { field_simp, ring, }, + rw [this, mul_lt_mul_left (one_half_pos : (0:ℝ) < 1/2)], + have : ‖∫ (x : ℝ), exp (↑(t * x) * I) • (f x - f (x + π / t))‖ ≤ ∫ (x : ℝ), + ‖exp (↑(t * x) * I) • (f x - f (x + π / t))‖, from norm_integral_le_integral_norm _, + refine lt_of_le_of_lt this _, + simp_rw [norm_smul, norm_exp_of_real_mul_I, one_mul], + -- Show integral can be taken over `[-(R + 1), R] ⊂ ℝ`. + let A := Icc (-(R + 1)) R, + have int_Icc : ∫ (x : ℝ), ‖f x - f (x + π / t)‖ = ∫ x in A, ‖f x - f (x + π / t)‖, + { refine (set_integral_eq_integral_of_forall_compl_eq_zero (λ x hx, _)).symm, + rw [mem_Icc, not_and_distrib, not_le, not_le, lt_neg] at hx, + suffices : (f x = 0 ∧ f (x + π / t) = 0), by { rw [this.1, this.2, sub_zero, norm_zero], }, + have tp : 0 < t := real.pi_pos.trans_le ((le_max_left _ _).trans ht), + refine ⟨hR2 x $ le_abs.mpr _, hR2 _ $ le_abs.mpr _⟩, + { cases hx, + { exact or.inr ((le_add_of_nonneg_right zero_le_one).trans hx.le) }, + { exact or.inl hx.le } }, + { cases hx, + { refine or.inr _, + rw [neg_add, ←sub_eq_add_neg, le_sub_iff_add_le], + refine le_trans (add_le_add_left _ R) hx.le, + exact (div_le_one tp).mpr ((le_max_left _ _).trans ht) }, + { exact or.inl (hx.trans $ lt_add_of_pos_right _ $ div_pos real.pi_pos tp).le } } }, + rw int_Icc, + -- Bound integral using fact that ‖f x - f (x + π / t)‖ is small. + have bdA : ∀ x : ℝ, (x ∈ A) → ‖ ‖f x - f (x + π / t) ‖ ‖ ≤ ε / (1 + 2 * R), + { simp_rw norm_norm, + refine (λ x _, le_of_lt _), + simp_rw dist_eq_norm at hδ2, + apply hδ2, + rw [sub_add_cancel', real.norm_eq_abs, abs_neg, abs_of_pos (div_pos real.pi_pos tpos), + div_lt_iff tpos, mul_comm, ←div_lt_iff hδ1], + linarith [(le_max_right π (1 + π / δ)).trans ht] }, + have bdA2 := norm_set_integral_le_of_norm_le_const (measure_Icc_lt_top : volume A < ∞) bdA _, + swap, { apply continuous.ae_strongly_measurable, + exact (continuous_norm.comp $ continuous.sub hf1 $ continuous.comp hf1 $ + continuous_id'.add continuous_const) }, + have : ‖ _ ‖ = ∫ (x : ℝ) in A, ‖f x - f (x + π / t)‖ := + real.norm_of_nonneg (set_integral_nonneg measurable_set_Icc (λ x hx, norm_nonneg _)), + rw this at bdA2, + refine lt_of_le_of_lt bdA2 _, + rw [real.volume_Icc, (by ring : R - (-(R + 1)) = 1 + 2 * R)], + have hh : 0 < 1 + 2 * R, by positivity, + rw [ennreal.to_real_of_real hh.le, div_mul_cancel _ hh.ne', two_mul], + exact lt_add_of_pos_left _ hε, +end + +end continuous_compact_support diff --git a/src/analysis/normed/group/basic.lean b/src/analysis/normed/group/basic.lean index 0579e29219a64..075c92a279765 100644 --- a/src/analysis/normed/group/basic.lean +++ b/src/analysis/normed/group/basic.lean @@ -1387,6 +1387,23 @@ by simpa [bdd_above_def] using hf.norm.bdd_above_range_of_has_compact_support h. end normed_add_group +section normed_add_group_source + +variables [normed_add_group α] {f : α → E} + +@[to_additive] +lemma has_compact_mul_support.exists_pos_le_norm [has_one E] (hf : has_compact_mul_support f) : + ∃ (R : ℝ), (0 < R) ∧ (∀ (x : α), (R ≤ ‖x‖) → (f x = 1)) := +begin + obtain ⟨K, ⟨hK1, hK2⟩⟩ := exists_compact_iff_has_compact_mul_support.mpr hf, + obtain ⟨S, hS, hS'⟩ := hK1.bounded.exists_pos_norm_le, + refine ⟨S + 1, by positivity, λ x hx, hK2 x ((mt $ hS' x) _)⟩, + contrapose! hx, + exact lt_add_of_le_of_pos hx zero_lt_one +end + +end normed_add_group_source + /-! ### `ulift` -/ namespace ulift diff --git a/src/measure_theory/integral/set_integral.lean b/src/measure_theory/integral/set_integral.lean index 85c9c82763295..02b274b9bf2cb 100644 --- a/src/measure_theory/integral/set_integral.lean +++ b/src/measure_theory/integral/set_integral.lean @@ -388,6 +388,23 @@ lemma set_integral_eq_of_subset_of_forall_diff_eq_zero set_integral_eq_of_subset_of_ae_diff_eq_zero ht.null_measurable_set hts (eventually_of_forall (λ x hx, h't x hx)) +/-- If a function vanishes almost everywhere on `sᶜ`, then its integral on `s` +coincides with its integral on the whole space. -/ +lemma set_integral_eq_integral_of_ae_compl_eq_zero + (h : ∀ᵐ x ∂μ, x ∉ s → f x = 0) : ∫ x in s, f x ∂μ = ∫ x, f x ∂μ := +begin + conv_rhs { rw ← integral_univ }, + symmetry, + apply set_integral_eq_of_subset_of_ae_diff_eq_zero null_measurable_set_univ (subset_univ _), + filter_upwards [h] with x hx h'x using hx h'x.2, +end + +/-- If a function vanishes on `sᶜ`, then its integral on `s` coincides with its integral on the +whole space. -/ +lemma set_integral_eq_integral_of_forall_compl_eq_zero + (h : ∀ x, x ∉ s → f x = 0) : ∫ x in s, f x ∂μ = ∫ x, f x ∂μ := +set_integral_eq_integral_of_ae_compl_eq_zero (eventually_of_forall h) + lemma set_integral_neg_eq_set_integral_nonpos [linear_order E] {f : α → E} (hf : ae_strongly_measurable f μ) : ∫ x in {x | f x < 0}, f x ∂μ = ∫ x in {x | f x ≤ 0}, f x ∂μ := diff --git a/src/topology/continuous_function/zero_at_infty.lean b/src/topology/continuous_function/zero_at_infty.lean index be667aa1b3767..aa34377be2ed4 100644 --- a/src/topology/continuous_function/zero_at_infty.lean +++ b/src/topology/continuous_function/zero_at_infty.lean @@ -291,7 +291,7 @@ variables [uniform_space β] [uniform_space γ] [has_zero γ] [zero_at_infty_continuous_map_class F β γ] lemma uniform_continuous (f : F) : uniform_continuous (f : β → γ) := -(map_continuous f).uniform_continuous_of_zero_at_infty (zero_at_infty f) +(map_continuous f).uniform_continuous_of_tendsto_cocompact (zero_at_infty f) end uniform diff --git a/src/topology/uniform_space/compact.lean b/src/topology/uniform_space/compact.lean index bb21c542896f3..21b9ad00f8784 100644 --- a/src/topology/uniform_space/compact.lean +++ b/src/topology/uniform_space/compact.lean @@ -6,6 +6,7 @@ Authors: Patrick Massot, Yury Kudryashov import topology.uniform_space.uniform_convergence import topology.uniform_space.equicontinuity import topology.separation +import topology.support /-! # Compact separated uniform spaces @@ -197,20 +198,40 @@ begin exacts [mem_ball_self _ (hT a a.2), mem_Inter₂.1 h a ha], end -lemma continuous.uniform_continuous_of_zero_at_infty {f : α → β} [has_zero β] - (h_cont : continuous f) (h_zero : tendsto f (cocompact α) (𝓝 0)) : uniform_continuous f := +lemma continuous.uniform_continuous_of_tendsto_cocompact {f : α → β} {x : β} + (h_cont : continuous f) (hx : tendsto f (cocompact α) (𝓝 x)) : uniform_continuous f := uniform_continuous_def.2 $ λ r hr, begin obtain ⟨t, ht, htsymm, htr⟩ := comp_symm_mem_uniformity_sets hr, - obtain ⟨s, hs, hst⟩ := mem_cocompact.1 (h_zero $ mem_nhds_left 0 ht), + obtain ⟨s, hs, hst⟩ := mem_cocompact.1 (hx $ mem_nhds_left _ ht), apply mem_of_superset (symmetrize_mem_uniformity $ hs.uniform_continuous_at_of_continuous_at f (λ _ _, h_cont.continuous_at) $ symmetrize_mem_uniformity hr), rintro ⟨b₁, b₂⟩ h, by_cases h₁ : b₁ ∈ s, { exact (h.1 h₁).1 }, by_cases h₂ : b₂ ∈ s, { exact (h.2 h₂).2 }, apply htr, - exact ⟨0, htsymm.mk_mem_comm.1 (hst h₁), hst h₂⟩, + exact ⟨x, htsymm.mk_mem_comm.1 (hst h₁), hst h₂⟩, end +/-- If `f` has compact multiplicative support, then `f` tends to 1 at infinity. -/ +@[to_additive "If `f` has compact support, then `f` tends to zero at infinity."] +lemma has_compact_mul_support.is_one_at_infty {f : α → γ} [topological_space γ] [has_one γ] + (h : has_compact_mul_support f) : tendsto f (cocompact α) (𝓝 1) := +begin + -- porting note: move to src/topology/support.lean once the port is over + intros N hN, + rw [mem_map, mem_cocompact'], + refine ⟨mul_tsupport f, h.is_compact, _⟩, + rw compl_subset_comm, + intros v hv, + rw [mem_preimage, image_eq_one_of_nmem_mul_tsupport hv], + exact mem_of_mem_nhds hN, +end + +@[to_additive] +lemma has_compact_mul_support.uniform_continuous_of_continuous {f : α → β} [has_one β] + (h1 : has_compact_mul_support f) (h2 : continuous f) : uniform_continuous f := +h2.uniform_continuous_of_tendsto_cocompact h1.is_one_at_infty + /-- A family of functions `α → β → γ` tends uniformly to its value at `x` if `α` is locally compact, `β` is compact and `f` is continuous on `U × (univ : set β)` for some neighborhood `U` of `x`. -/ lemma continuous_on.tendsto_uniformly [locally_compact_space α] [compact_space β] From d3e3adc967c909c350a088d378644b53f10c9f92 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 10 Feb 2023 12:21:17 +0000 Subject: [PATCH 0457/1291] chore(algebra/quaternion): missing trivial lemmas (#18413) A mixture of trivial algebraic results and trivial topological ones. This also changes the `rat.cast` instance on quaternions to be slightly nicer. --- src/algebra/quaternion.lean | 88 ++++++++++++++++++++++++++++++++++-- src/analysis/quaternion.lean | 20 ++++++++ 2 files changed, 104 insertions(+), 4 deletions(-) diff --git a/src/algebra/quaternion.lean b/src/algebra/quaternion.lean index 78174e14b61b6..8822b8292c4e3 100644 --- a/src/algebra/quaternion.lean +++ b/src/algebra/quaternion.lean @@ -174,6 +174,18 @@ instance : add_group_with_one ℍ[R, c₁, c₂] := one := 1, .. quaternion_algebra.add_comm_group } +@[simp, norm_cast] lemma nat_cast_re (n : ℕ) : (n : ℍ[R, c₁, c₂]).re = n := rfl +@[simp, norm_cast] lemma nat_cast_im_i (n : ℕ) : (n : ℍ[R, c₁, c₂]).im_i = 0 := rfl +@[simp, norm_cast] lemma nat_cast_im_j (n : ℕ) : (n : ℍ[R, c₁, c₂]).im_j = 0 := rfl +@[simp, norm_cast] lemma nat_cast_im_k (n : ℕ) : (n : ℍ[R, c₁, c₂]).im_k = 0 := rfl +@[norm_cast] lemma coe_nat_cast (n : ℕ) : ↑(n : R) = (n : ℍ[R, c₁, c₂]) := rfl + +@[simp, norm_cast] lemma int_cast_re (z : ℤ) : (z : ℍ[R, c₁, c₂]).re = z := rfl +@[simp, norm_cast] lemma int_cast_im_i (z : ℤ) : (z : ℍ[R, c₁, c₂]).im_i = 0 := rfl +@[simp, norm_cast] lemma int_cast_im_j (z : ℤ) : (z : ℍ[R, c₁, c₂]).im_j = 0 := rfl +@[simp, norm_cast] lemma int_cast_im_k (z : ℤ) : (z : ℍ[R, c₁, c₂]).im_k = 0 := rfl +@[norm_cast] lemma coe_int_cast (z : ℤ) : ↑(z : R) = (z : ℍ[R, c₁, c₂]) := rfl + instance : ring ℍ[R, c₁, c₂] := by refine_struct { add := (+), @@ -262,6 +274,9 @@ end @[norm_cast, simp] lemma coe_mul : ((x * y : R) : ℍ[R, c₁, c₂]) = x * y := (algebra_map R ℍ[R, c₁, c₂]).map_mul x y +@[norm_cast, simp] lemma coe_pow (n : ℕ) : (↑(x ^ n) : ℍ[R, c₁, c₂]) = ↑x ^ n := +(algebra_map R ℍ[R, c₁, c₂]).map_pow x n + lemma coe_commutes : ↑r * a = a * r := algebra.commutes r a lemma coe_commute : commute ↑r a := coe_commutes r a @@ -329,7 +344,12 @@ calc a.conj * b.conj = (b * a).conj : (conj_mul b a).symm ... = (a * b).conj : by rw h.eq ... = b.conj * a.conj : conj_mul a b -@[simp] lemma conj_coe : conj (x : ℍ[R, c₁, c₂]) = x := by ext; simp +@[simp, norm_cast] lemma conj_coe : conj (x : ℍ[R, c₁, c₂]) = x := by ext; simp + +@[simp, norm_cast] lemma conj_nat_cast (n : ℕ) : conj (n : ℍ[R, c₁, c₂]) = n := +by rw [←coe_nat_cast, conj_coe] +@[simp, norm_cast] lemma conj_int_cast (z : ℤ) : conj (z : ℍ[R, c₁, c₂]) = z := +by rw [←coe_int_cast, conj_coe] lemma conj_smul : conj (r • a) = r • conj a := conj.map_smul r a @@ -369,6 +389,8 @@ instance : star_ring ℍ[R, c₁, c₂] := @[simp] lemma star_def (a : ℍ[R, c₁, c₂]) : star a = conj a := rfl +@[simp] lemma conj_pow (n : ℕ) : (a ^ n).conj = a.conj ^ n := star_pow _ _ + open mul_opposite /-- Quaternion conjugate as an `alg_equiv` to the opposite ring. -/ @@ -478,6 +500,21 @@ quaternion_algebra.ext_iff a b @[simp, norm_cast] lemma coe_mul : ((x * y : R) : ℍ[R]) = x * y := quaternion_algebra.coe_mul x y +@[norm_cast, simp] lemma coe_pow (n : ℕ) : (↑(x ^ n) : ℍ[R]) = ↑x ^ n := +quaternion_algebra.coe_pow x n + +@[simp, norm_cast] lemma nat_cast_re (n : ℕ) : (n : ℍ[R]).re = n := rfl +@[simp, norm_cast] lemma nat_cast_im_i (n : ℕ) : (n : ℍ[R]).im_i = 0 := rfl +@[simp, norm_cast] lemma nat_cast_im_j (n : ℕ) : (n : ℍ[R]).im_j = 0 := rfl +@[simp, norm_cast] lemma nat_cast_im_k (n : ℕ) : (n : ℍ[R]).im_k = 0 := rfl +@[norm_cast] lemma coe_nat_cast (n : ℕ) : ↑(n : R) = (n : ℍ[R]) := rfl + +@[simp, norm_cast] lemma int_cast_re (z : ℤ) : (z : ℍ[R]).re = z := rfl +@[simp, norm_cast] lemma int_cast_im_i (z : ℤ) : (z : ℍ[R]).im_i = 0 := rfl +@[simp, norm_cast] lemma int_cast_im_j (z : ℤ) : (z : ℍ[R]).im_j = 0 := rfl +@[simp, norm_cast] lemma int_cast_im_k (z : ℤ) : (z : ℍ[R]).im_k = 0 := rfl +@[norm_cast] lemma coe_int_cast (z : ℤ) : ↑(z : R) = (z : ℍ[R]) := rfl + lemma coe_injective : function.injective (coe : R → ℍ[R]) := quaternion_algebra.coe_injective @[simp] lemma coe_inj {x y : R} : (x : ℍ[R]) = y ↔ x = y := coe_injective.eq_iff @@ -545,7 +582,12 @@ quaternion_algebra.commute_conj_conj h alias commute_conj_conj ← commute.quaternion_conj -@[simp] lemma conj_coe : conj (x : ℍ[R]) = x := quaternion_algebra.conj_coe x +@[simp, norm_cast] lemma conj_coe : conj (x : ℍ[R]) = x := quaternion_algebra.conj_coe x + +@[simp, norm_cast] lemma conj_nat_cast (n : ℕ) : conj (n : ℍ[R]) = n := +quaternion_algebra.conj_nat_cast _ +@[simp, norm_cast] lemma conj_int_cast (z : ℤ) : conj (z : ℍ[R]) = z := +quaternion_algebra.conj_int_cast _ @[simp] lemma conj_smul : conj (r • a) = r • conj a := a.conj_smul r @@ -571,6 +613,8 @@ lemma mul_conj_eq_coe : a * conj a = (a * conj a).re := a.mul_conj_eq_coe @[simp] lemma conj_sub : (a - b).conj = a.conj - b.conj := a.conj_sub b +@[simp] lemma conj_pow (n : ℕ) : conj (a ^ n) = conj a ^ n := a.conj_pow n + open mul_opposite /-- Quaternion conjugate as an `alg_equiv` to the opposite ring. -/ @@ -595,6 +639,14 @@ by simp only [norm_sq_def, sq, mul_neg, sub_neg_eq_add, lemma norm_sq_coe : norm_sq (x : ℍ[R]) = x^2 := by rw [norm_sq_def, conj_coe, ← coe_mul, coe_re, sq] +@[simp] lemma norm_sq_conj : norm_sq (conj a) = norm_sq a := by simp [norm_sq_def'] + +@[norm_cast] lemma norm_sq_nat_cast (n : ℕ) : norm_sq (n : ℍ[R]) = n^2 := +by rw [←coe_nat_cast, norm_sq_coe] + +@[norm_cast] lemma norm_sq_int_cast (z : ℤ) : norm_sq (z : ℍ[R]) = z^2 := +by rw [←coe_int_cast, norm_sq_coe] + @[simp] lemma norm_sq_neg : norm_sq (-a) = norm_sq a := by simp only [norm_sq_def, conj_neg, neg_mul_neg] @@ -650,18 +702,46 @@ section field variables [linear_ordered_field R] (a b : ℍ[R]) -@[simps { attrs := [] }]instance : has_inv ℍ[R] := ⟨λ a, (norm_sq a)⁻¹ • a.conj⟩ +@[simps { attrs := [] }] instance : has_inv ℍ[R] := ⟨λ a, (norm_sq a)⁻¹ • a.conj⟩ -instance : division_ring ℍ[R] := +instance : group_with_zero ℍ[R] := { inv := has_inv.inv, inv_zero := by rw [has_inv_inv, conj_zero, smul_zero], mul_inv_cancel := λ a ha, by rw [has_inv_inv, algebra.mul_smul_comm, self_mul_conj, smul_coe, inv_mul_cancel (norm_sq_ne_zero.2 ha), coe_one], .. quaternion.nontrivial, + .. (by apply_instance : monoid_with_zero ℍ[R]) } + +@[norm_cast, simp] lemma coe_inv (x : R) : ((x⁻¹ : R) : ℍ[R]) = x⁻¹ := +map_inv₀ (algebra_map R ℍ[R]) _ + +@[norm_cast, simp] lemma coe_div (x y : R) : ((x / y : R) : ℍ[R]) = x / y := +map_div₀ (algebra_map R ℍ[R]) x y + +@[norm_cast, simp] lemma coe_zpow (x : R) (z : ℤ) : ((x ^ z : R) : ℍ[R]) = x ^ z := +map_zpow₀ (algebra_map R ℍ[R]) x z + +instance : division_ring ℍ[R] := +{ rat_cast := λ q, ↑(q : R), + rat_cast_mk := λ n d hd h, by rw [rat.cast_mk', coe_mul, coe_int_cast, coe_inv, coe_nat_cast], + .. quaternion.group_with_zero, .. quaternion.ring } +@[simp, norm_cast] lemma rat_cast_re (q : ℚ) : (q : ℍ[R]).re = q := rfl +@[simp, norm_cast] lemma rat_cast_im_i (q : ℚ) : (q : ℍ[R]).im_i = 0 := rfl +@[simp, norm_cast] lemma rat_cast_im_j (q : ℚ) : (q : ℍ[R]).im_j = 0 := rfl +@[simp, norm_cast] lemma rat_cast_im_k (q : ℚ) : (q : ℍ[R]).im_k = 0 := rfl +@[norm_cast] lemma coe_rat_cast (q : ℚ) : ↑(q : R) = (q : ℍ[R]) := rfl + +lemma conj_inv : conj (a⁻¹) = (conj a)⁻¹ := star_inv' a +lemma conj_zpow (z : ℤ) : conj (a ^ z) = conj a ^ z := star_zpow₀ a z +@[simp, norm_cast] lemma conj_rat_cast (q : ℚ) : conj (q : ℍ[R]) = q := @star_rat_cast ℍ[R] _ _ q + @[simp] lemma norm_sq_inv : norm_sq a⁻¹ = (norm_sq a)⁻¹ := map_inv₀ norm_sq _ @[simp] lemma norm_sq_div : norm_sq (a / b) = norm_sq a / norm_sq b := map_div₀ norm_sq a b +@[simp] lemma norm_sq_zpow (z : ℤ) : norm_sq (a ^ z) = norm_sq a ^ z := map_zpow₀ norm_sq a z +@[norm_cast] lemma norm_sq_rat_cast (q : ℚ) : norm_sq (q : ℍ[R]) = q^2 := +by rw [←coe_rat_cast, norm_sq_coe] end field diff --git a/src/analysis/quaternion.lean b/src/analysis/quaternion.lean index 4fd74563713f3..b5c499aa8dd2a 100644 --- a/src/analysis/quaternion.lean +++ b/src/analysis/quaternion.lean @@ -6,6 +6,7 @@ Authors: Yury Kudryashov, Eric Wieser import algebra.quaternion import analysis.inner_product_space.basic import analysis.inner_product_space.pi_L2 +import topology.algebra.algebra /-! # Quaternions as a normed algebra @@ -61,6 +62,12 @@ by rw [norm_eq_sqrt_real_inner, inner_self, norm_sq_coe, real.sqrt_sq_eq_abs, re @[simp, norm_cast] lemma nnnorm_coe (a : ℝ) : ‖(a : ℍ)‖₊ = ‖a‖₊ := subtype.ext $ norm_coe a +@[simp] lemma norm_conj (a : ℍ) : ‖conj a‖ = ‖a‖ := +by simp_rw [norm_eq_sqrt_real_inner, inner_self, norm_sq_conj] + +@[simp] lemma nnnorm_conj (a : ℍ) : ‖conj a‖₊ = ‖a‖₊ := +subtype.ext $ norm_conj a + noncomputable instance : normed_division_ring ℍ := { dist_eq := λ _ _, rfl, norm_mul' := λ a b, by { simp only [norm_eq_sqrt_real_inner, inner_self, norm_sq.map_mul], @@ -70,6 +77,9 @@ instance : normed_algebra ℝ ℍ := { norm_smul_le := λ a x, (norm_smul a x).le, to_algebra := quaternion.algebra } +instance : cstar_ring ℍ := +{ norm_star_mul_self := λ x, (norm_mul _ _).trans $ congr_arg (* ‖x‖) (norm_conj x) } + instance : has_coe ℂ ℍ := ⟨λ z, ⟨z.re, z.im, 0, 0⟩⟩ @[simp, norm_cast] lemma coe_complex_re (z : ℂ) : (z : ℍ).re = z.re := rfl @@ -115,6 +125,16 @@ noncomputable def linear_isometry_equiv_tuple : ℍ ≃ₗᵢ[ℝ] euclidean_spa ..(quaternion_algebra.linear_equiv_tuple (-1 : ℝ) (-1 : ℝ)).trans (pi_Lp.linear_equiv 2 ℝ (λ _ : fin 4, ℝ)).symm } +@[continuity] lemma continuous_conj : continuous (conj : ℍ → ℍ) := +continuous_star + +@[continuity] lemma continuous_coe : continuous (coe : ℝ → ℍ) := +continuous_algebra_map ℝ ℍ + +@[continuity] lemma continuous_norm_sq : continuous (norm_sq : ℍ → ℝ) := +by simpa [←norm_sq_eq_norm_sq] + using (continuous_norm.mul continuous_norm : continuous (λ q : ℍ, ‖q‖ * ‖q‖)) + @[continuity] lemma continuous_re : continuous (λ q : ℍ, q.re) := (continuous_apply 0).comp linear_isometry_equiv_tuple.continuous From 0f1becb755b3d008b242c622e248a70556ad19e6 Mon Sep 17 00:00:00 2001 From: sgouezel Date: Fri, 10 Feb 2023 13:29:34 +0000 Subject: [PATCH 0458/1291] chore(measure_theory/decomposition/unsigned_hahn): speedup Hahn decomposition (#18419) From 49s to 4s on my computer. The key point was to replace the super-slow `ac_refl` with `abel`. --- .../decomposition/unsigned_hahn.lean | 16 ++++++---------- 1 file changed, 6 insertions(+), 10 deletions(-) diff --git a/src/measure_theory/decomposition/unsigned_hahn.lean b/src/measure_theory/decomposition/unsigned_hahn.lean index 998240ec6148a..612c2f03d4633 100644 --- a/src/measure_theory/decomposition/unsigned_hahn.lean +++ b/src/measure_theory/decomposition/unsigned_hahn.lean @@ -28,11 +28,6 @@ namespace measure_theory variables {α : Type*} [measurable_space α] {μ ν : measure α} --- suddenly this is necessary?! -private lemma aux {m : ℕ} {γ d : ℝ} (h : γ - (1 / 2) ^ m < d) : - γ - 2 * (1 / 2) ^ m + (1 / 2) ^ m ≤ d := -by linarith - /-- **Hahn decomposition theorem** -/ lemma hahn_decomposition [is_finite_measure μ] [is_finite_measure ν] : ∃s, measurable_set s ∧ @@ -61,7 +56,7 @@ begin ennreal.to_nnreal_add (hμ _) (hμ _), ennreal.to_nnreal_add (hν _) (hν _), nnreal.coe_add, nnreal.coe_add], simp only [sub_eq_add_neg, neg_add], - ac_refl }, + abel }, have d_Union : ∀(s : ℕ → set α), monotone s → tendsto (λn, d (s n)) at_top (𝓝 (d (⋃n, s n))), @@ -127,7 +122,7 @@ begin { have := he₂ m, simp only [f], rw [nat.Ico_succ_singleton, finset.inf_singleton], - exact aux this }, + linarith }, { assume n (hmn : m ≤ n) ih, have : γ + (γ - 2 * (1 / 2)^m + (1 / 2) ^ (n + 1)) ≤ γ + d (f m (n + 1)), { calc γ + (γ - 2 * (1 / 2)^m + (1 / 2) ^ (n+1)) ≤ @@ -138,7 +133,7 @@ begin linarith end ... = (γ - (1 / 2)^(n+1)) + (γ - 2 * (1 / 2)^m + (1 / 2)^n) : - by simp only [sub_eq_add_neg]; ac_refl + by simp only [sub_eq_add_neg]; abel ... ≤ d (e (n + 1)) + d (f m n) : add_le_add (le_of_lt $ he₂ _) ih ... ≤ d (e (n + 1)) + d (f m n \ e (n + 1)) + d (f m (n + 1)) : by rw [f_succ _ _ hmn, d_split (f m n) (e (n + 1)) (hf _ _) (he₁ _), add_assoc] @@ -146,7 +141,7 @@ begin begin rw [d_split (e (n + 1) ∪ f m n) (e (n + 1)), union_diff_left, union_inter_cancel_left], - ac_refl, + abel, exact (he₁ _).union (hf _ _), exact (he₁ _) end @@ -157,7 +152,8 @@ begin let s := ⋃ m, ⋂n, f m n, have γ_le_d_s : γ ≤ d s, { have hγ : tendsto (λm:ℕ, γ - 2 * (1/2)^m) at_top (𝓝 γ), - { suffices : tendsto (λm:ℕ, γ - 2 * (1/2)^m) at_top (𝓝 (γ - 2 * 0)), { simpa }, + { suffices : tendsto (λm:ℕ, γ - 2 * (1/2)^m) at_top (𝓝 (γ - 2 * 0)), + { simpa only [mul_zero, tsub_zero] }, exact (tendsto_const_nhds.sub $ tendsto_const_nhds.mul $ tendsto_pow_at_top_nhds_0_of_lt_1 (le_of_lt $ half_pos $ zero_lt_one) (half_lt_self zero_lt_one)) }, From 4b83fd2855cbf5e9da1fcf0fc1eba89a48c21551 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Fri, 10 Feb 2023 14:36:05 +0000 Subject: [PATCH 0459/1291] feat(measure_theory/group.measure): add pi.is_haar_measure (#18390) Add the proof that a finite product of Haar measures is a Haar measure by proving the following instances `pi.is_open_pos_measure` and `pi.is_finite_measure_on_compacts` It is also necessary to prove a lemma `is_open_pi_iff` belonging to an already ported file. The corresponding mathlib4 PR is: https://github.com/leanprover-community/mathlib4/pull/2110 Co-authored-by: Eric Wieser --- src/measure_theory/constructions/pi.lean | 32 +++++++++++++++++ src/topology/constructions.lean | 45 ++++++++++++++++++++++++ 2 files changed, 77 insertions(+) diff --git a/src/measure_theory/constructions/pi.lean b/src/measure_theory/constructions/pi.lean index b2c120dbd7656..4fb749a467fee 100644 --- a/src/measure_theory/constructions/pi.lean +++ b/src/measure_theory/constructions/pi.lean @@ -5,6 +5,7 @@ Authors: Floris van Doorn -/ import measure_theory.constructions.prod import measure_theory.group.measure +import topology.constructions /-! # Product measures @@ -538,6 +539,37 @@ begin pi_pi, measure_preimage_inv] end +instance pi.is_open_pos_measure [Π i, topological_space (α i)] [Π i, is_open_pos_measure (μ i)] : + is_open_pos_measure (measure_theory.measure.pi μ) := +begin + constructor, + rintros U U_open ⟨a, ha⟩, + obtain ⟨s, ⟨hs, hsU⟩⟩ := is_open_pi_iff'.1 U_open a ha, + refine ne_of_gt (lt_of_lt_of_le _ (measure_mono hsU)), + simp only [pi_pi], + rw canonically_ordered_comm_semiring.prod_pos, + intros i _, + apply ((hs i).1.measure_pos (μ i) ⟨a i, (hs i).2⟩), +end + +instance pi.is_finite_measure_on_compacts [Π i, topological_space (α i)] + [Π i, is_finite_measure_on_compacts (μ i)] : + is_finite_measure_on_compacts (measure_theory.measure.pi μ) := +begin + constructor, + intros K hK, + suffices : measure.pi μ (set.univ.pi ( λ j, (function.eval j) '' K)) < ⊤, + { exact lt_of_le_of_lt (measure_mono (univ.subset_pi_eval_image K)) this, }, + rw measure.pi_pi, + refine with_top.prod_lt_top _, + exact λ i _, ne_of_lt (is_compact.measure_lt_top (is_compact.image hK (continuous_apply i))), +end + +@[to_additive] +instance pi.is_haar_measure [Π i, group (α i)] [Π i, topological_space (α i)] + [Π i, is_haar_measure (μ i)] [Π i, has_measurable_mul (α i)] : + is_haar_measure (measure.pi μ) := {} + end measure instance measure_space.pi [Π i, measure_space (α i)] : measure_space (Π i, α i) := ⟨measure.pi (λ i, volume)⟩ diff --git a/src/topology/constructions.lean b/src/topology/constructions.lean index 9de0f22983304..ad136bd9b80b4 100644 --- a/src/topology/constructions.lean +++ b/src/topology/constructions.lean @@ -1061,6 +1061,51 @@ lemma is_open_set_pi {i : set ι} {s : Πa, set (π a)} (hi : i.finite) (hs : is_open (pi i s) := by rw [pi_def]; exact (is_open_bInter hi $ assume a ha, (hs _ ha).preimage (continuous_apply _)) +lemma is_open_pi_iff {s : set (Π a, π a)} : + is_open s ↔ + (∀ f, f ∈ s → ∃ (I : finset ι) (u : Π a, set (π a)), + (∀ a, a ∈ I → is_open (u a) ∧ f a ∈ u a) ∧ (I : set ι).pi u ⊆ s) := +begin + rw is_open_iff_nhds, + simp_rw [le_principal_iff, nhds_pi, filter.mem_pi', mem_nhds_iff, exists_prop], + refine ball_congr (λ a h, ⟨_, _⟩), + { rintros ⟨I, t, ⟨h1, h2⟩⟩, + refine ⟨I, λ a, eval a '' ((I : set ι).pi (λ a, (h1 a).some)), (λ i hi, _), _⟩, + { simp_rw set.eval_image_pi (finset.mem_coe.mpr hi) + (pi_nonempty_iff.mpr (λ i, ⟨_, λ _, (h1 i).some_spec.2.2⟩)), + exact (h1 i).some_spec.2, }, + { refine subset.trans + (set.pi_mono (λ i hi, (set.eval_image_pi_subset hi).trans (h1 i).some_spec.1)) h2, }}, + { rintros ⟨I, t, ⟨h1, h2⟩⟩, + refine ⟨I, λ a, ite (a ∈ I) (t a) (set.univ), (λ i, _), _⟩, + { by_cases hi : i ∈ I, + { use t i, + rw if_pos hi, + exact ⟨subset.rfl, (h1 i) hi⟩, }, + { use set.univ, + rw if_neg hi, + exact ⟨subset.rfl, is_open_univ, mem_univ _⟩, }}, + { rw ← set.univ_pi_ite, + simp only [ ← ite_and, ← finset.mem_coe, and_self, set.univ_pi_ite, h2], }} +end + +lemma is_open_pi_iff' [finite ι] {s : set (Π a, π a)} : + is_open s ↔ + (∀ f, f ∈ s → ∃ (u : Π a, set (π a)), (∀ a, is_open (u a) ∧ f a ∈ u a) ∧ set.univ.pi u ⊆ s) := +begin + casesI nonempty_fintype ι, + rw is_open_iff_nhds, + simp_rw [le_principal_iff, nhds_pi, filter.mem_pi', mem_nhds_iff, exists_prop], + refine ball_congr (λ a h, ⟨_, _⟩), + { rintros ⟨I, t, ⟨h1, h2⟩⟩, + refine ⟨λ i, (h1 i).some, ⟨λ i, (h1 i).some_spec.2, + (set.pi_mono (λ i _, (h1 i).some_spec.1)).trans (subset.trans _ h2)⟩⟩, + rw ← set.pi_inter_compl (I : set ι), + exact inter_subset_left _ _, }, + { exact λ ⟨u, ⟨h1, _⟩⟩, ⟨finset.univ, u, ⟨λ i, ⟨u i, ⟨rfl.subset, h1 i⟩⟩, + by rwa finset.coe_univ⟩⟩, } +end + lemma is_closed_set_pi {i : set ι} {s : Πa, set (π a)} (hs : ∀a∈i, is_closed (s a)) : is_closed (pi i s) := by rw [pi_def]; From 67f92b6f60865717e0811128b77d49ab4c07f7f1 Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Fri, 10 Feb 2023 16:39:37 +0000 Subject: [PATCH 0460/1291] feat(algebraic_geometry/elliptic_curve/point): define addition of K-rational points (#17194) - [x] depends on: #17631 - [x] depends on: #17700 - [x] depends on: #17911 Co-authored-by: Junyan Xu --- .../elliptic_curve/point.lean | 473 ++++++++++++++++++ .../elliptic_curve/weierstrass.lean | 112 +++-- 2 files changed, 534 insertions(+), 51 deletions(-) create mode 100644 src/algebraic_geometry/elliptic_curve/point.lean diff --git a/src/algebraic_geometry/elliptic_curve/point.lean b/src/algebraic_geometry/elliptic_curve/point.lean new file mode 100644 index 0000000000000..7079473cef4fd --- /dev/null +++ b/src/algebraic_geometry/elliptic_curve/point.lean @@ -0,0 +1,473 @@ +/- +Copyright (c) 2022 David Kurniadi Angdinata. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: David Kurniadi Angdinata +-/ + +import algebraic_geometry.elliptic_curve.weierstrass +import field_theory.galois -- temporary import to enable point notation +import ring_theory.class_group + +/-! +# The group of nonsingular rational points on a Weierstrass curve over a field + +This file defines the type of nonsingular rational points on a Weierstrass curve over a field and +(TODO) proves that it forms an abelian group under a geometric secant-and-tangent process. + +## Mathematical background + +Let `W` be a Weierstrass curve over a field `F`. A rational point on `W` is simply a point +$[A:B:C]$ defined over `F` in the projective plane satisfying the homogeneous cubic equation +$B^2C + a_1ABC + a_3BC^2 = A^3 + a_2A^2C + a_4AC^2 + a_6C^3$. Any such point either lies in the +affine chart $C \ne 0$ and satisfies the Weierstrass equation obtained by setting $X := A/C$ and +$Y := B/C$, or is the unique point at infinity $0 := [0:1:0]$ when $C = 0$. With this new +description, a nonsingular rational point on `W` is either $0$ or an affine point $(x, y)$ where +the partial derivatives $W_X(X, Y)$ and $W_Y(X, Y)$ do not vanish simultaneously. For a field +extension `K` of `F`, a `K`-rational point is simply a rational point on `W` base changed to `K`. + +The set of nonsingular rational points forms an abelian group under a secant-and-tangent process. + * The identity rational point is `0`. + * Given a nonsingular rational point `P`, its negation `-P` is defined to be the unique third + point of intersection between `W` and the line through `0` and `P`. + Explicitly, if `P` is $(x, y)$, then `-P` is $(x, -y - a_1x - a_3)$. + * Given two points `P` and `Q`, their addition `P + Q` is defined to be the negation of the unique + third point of intersection between `W` and the line `L` through `P` and `Q`. + Explicitly, let `P` be $(x_1, y_1)$ and let `Q` be $(x_2, y_2)$. + * If $x_1 = x_2$ and $y_1 = -y_2 - a_1x_2 - a_3$, then `L` is vertical and `P + Q` is `0`. + * If $x_1 = x_2$ and $y_1 \ne -y_2 - a_1x_2 - a_3$, then `L` is the tangent of `W` at `P = Q`, + and has slope $\ell := (3x_1^2 + 2a_2x_1 + a_4 - a_1y_1) / (2y_1 + a_1x_1 + a_3)$. + * Otherwise $x_1 \ne x_2$, then `L` is the secant of `W` through `P` and `Q`, and has slope + $\ell := (y_1 - y_2) / (x_1 - x_2)$. + In the latter two cases, the $X$-coordinate of `P + Q` is then the unique third solution of the + equation obtained by substituting the line $Y = \ell(X - x_1) + y_1$ into the Weierstrass + equation, and can be written down explicitly as $x := \ell^2 + a_1\ell - a_2 - x_1 - x_2$ by + inspecting the $X^2$ terms. The $Y$-coordinate of `P + Q`, after applying the final negation + that maps $Y$ to $-Y - a_1X - a_3$, is precisely $y := -(\ell(x - x_1) + y_1) - a_1x - a_3$. +The group law on this set is then uniquely determined by these constructions. + +## Main definitions + + * `weierstrass_curve.point`: the type of nonsingular rational points on a Weierstrass curve `W`. + * `weierstrass_curve.point.add`: the addition of two nonsingular rational points on `W`. + +## Main statements + + * TODO: the addition of two nonsingular rational points on `W` forms a group. + +## Notations + + * `W⟮K⟯`: the group of nonsingular rational points on a Weierstrass curve `W` base changed to `K`. + +## References + +[J Silverman, *The Arithmetic of Elliptic Curves*][silverman2009] + +## Tags + +elliptic curve, rational point, group law +-/ + +private meta def eval_simp : tactic unit := +`[simp only [eval_C, eval_X, eval_neg, eval_add, eval_sub, eval_mul, eval_pow]] + +private meta def C_simp : tactic unit := +`[simp only [C_1, C_bit0, C_bit1, C_neg, C_add, C_sub, C_mul, C_pow]] + +private meta def derivative_simp : tactic unit := +`[simp only [derivative_C, derivative_X, derivative_X_pow, derivative_neg, derivative_add, + derivative_sub, derivative_mul, derivative_sq]] + +universe u + +namespace weierstrass_curve + +open polynomial + +open_locale polynomial polynomial_polynomial + +section basic + +/-! ### Polynomials associated to nonsingular rational points on a Weierstrass curve -/ + +variables {R : Type u} [comm_ring R] (W : weierstrass_curve R) (x₁ x₂ y₁ y₂ L : R) + +/-- The polynomial $-Y - a_1X - a_3$ associated to negation. -/ +noncomputable def neg_polynomial : R[X][Y] := -Y - C (C W.a₁ * X + C W.a₃) + +/-- The $Y$-coordinate of the negation of an affine point in `W`. + +This depends on `W`, and has argument order: $x_1$, $y_1$. -/ +@[simp] def neg_Y : R := -y₁ - W.a₁ * x₁ - W.a₃ + +lemma neg_Y_neg_Y : W.neg_Y x₁ (W.neg_Y x₁ y₁) = y₁ := by { simp only [neg_Y], ring1 } + +@[simp] lemma eval_neg_polynomial : (W.neg_polynomial.eval $ C y₁).eval x₁ = W.neg_Y x₁ y₁ := +by { rw [neg_Y, sub_sub, neg_polynomial], eval_simp } + +/-- The polynomial $L*(X - x_1) + y_1$ associated to the line $Y = L*(X - x_1) + y_1$, +with a slope of $L$ that passes through an affine point $(x_1, y_1)$. + +This does not depend on `W`, and has argument order: $x_1$, $y_1$, $L$. -/ +noncomputable def line_polynomial : R[X] := C L * (X - C x₁) + C y₁ + +/-- The polynomial obtained by substituting the line $Y = L*(X - x_1) + y_1$, with a slope of $L$ +that passes through an affine point $(x_1, y_1)$, into the polynomial $W(X, Y)$ associated to `W`. +If such a line intersects `W` at another point $(x_2, y_2)$, then the roots of this polynomial are +precisely $x_1$, $x_2$, and the $X$-coordinate of the addition of $(x_1, y_1)$ and $(x_2, y_2)$. + +This depends on `W`, and has argument order: $x_1$, $y_1$, $L$. -/ +noncomputable def add_polynomial : R[X] := W.polynomial.eval $ line_polynomial x₁ y₁ L + +lemma add_polynomial_eq : W.add_polynomial x₁ y₁ L = -cubic.to_poly + ⟨1, -L ^ 2 - W.a₁ * L + W.a₂, + 2 * x₁ * L ^ 2 + (W.a₁ * x₁ - 2 * y₁ - W.a₃) * L + (-W.a₁ * y₁ + W.a₄), + -x₁ ^ 2 * L ^ 2 + (2 * x₁ * y₁ + W.a₃ * x₁) * L - (y₁ ^ 2 + W.a₃ * y₁ - W.a₆)⟩ := +by { rw [add_polynomial, line_polynomial, weierstrass_curve.polynomial, cubic.to_poly], eval_simp, + C_simp, ring1 } + +/-- The $X$-coordinate of the addition of two affine points $(x_1, y_1)$ and $(x_2, y_2)$ in `W`, +where the line through them is not vertical and has a slope of $L$. + +This depends on `W`, and has argument order: $x_1$, $x_2$, $L$. -/ +@[simp] def add_X : R := L ^ 2 + W.a₁ * L - W.a₂ - x₁ - x₂ + +/-- The $Y$-coordinate, before applying the final negation, of the addition of two affine points +$(x_1, y_1)$ and $(x_2, y_2)$, where the line through them is not vertical and has a slope of $L$. + +This depends on `W`, and has argument order: $x_1$, $x_2$, $y_1$, $L$. -/ +@[simp] def add_Y' : R := L * (W.add_X x₁ x₂ L - x₁) + y₁ + +/-- The $Y$-coordinate of the addition of two affine points $(x_1, y_1)$ and $(x_2, y_2)$ in `W`, +where the line through them is not vertical and has a slope of $L$. + +This depends on `W`, and has argument order: $x_1$, $x_2$, $y_1$, $L$. -/ +@[simp] def add_Y : R := W.neg_Y (W.add_X x₁ x₂ L) (W.add_Y' x₁ x₂ y₁ L) + +lemma equation_add_iff : + W.equation (W.add_X x₁ x₂ L) (W.add_Y' x₁ x₂ y₁ L) + ↔ (W.add_polynomial x₁ y₁ L).eval (W.add_X x₁ x₂ L) = 0 := +by { rw [equation, add_Y', add_polynomial, line_polynomial, weierstrass_curve.polynomial], + eval_simp } + +lemma nonsingular_add_of_eval_derivative_ne_zero + (hx' : W.equation (W.add_X x₁ x₂ L) (W.add_Y' x₁ x₂ y₁ L)) + (hx : (derivative $ W.add_polynomial x₁ y₁ L).eval (W.add_X x₁ x₂ L) ≠ 0) : + W.nonsingular (W.add_X x₁ x₂ L) (W.add_Y' x₁ x₂ y₁ L) := +begin + rw [nonsingular, and_iff_right hx', add_Y', polynomial_X, polynomial_Y], + eval_simp, + contrapose! hx, + rw [add_polynomial, line_polynomial, weierstrass_curve.polynomial], + eval_simp, + derivative_simp, + simp only [zero_add, add_zero, sub_zero, zero_mul, mul_one], + eval_simp, + linear_combination hx.left + L * hx.right with { normalization_tactic := `[norm_num1, ring1] } +end + +/-! ### The type of nonsingular rational points on a Weierstrass curve -/ + +/-- A nonsingular rational point on a Weierstrass curve `W` over `R`. This is either the point at +infinity `weierstrass_curve.point.zero` or an affine point `weierstrass_curve.point.some` $(x, y)$ +satisfying the equation $y^2 + a_1xy + a_3y = x^3 + a_2x^2 + a_4x + a_6$ of `W`. For an algebraic +extension `S` of `R`, the type of nonsingular `S`-rational points on `W` is denoted `W⟮S⟯`. -/ +inductive point +| zero +| some {x y : R} (h : W.nonsingular x y) + +localized "notation W⟮S⟯ := (W.base_change S).point" in weierstrass_curve + +namespace point + +instance : inhabited W.point := ⟨zero⟩ + +instance : has_zero W.point := ⟨zero⟩ + +@[simp] lemma zero_def : (zero : W.point) = 0 := rfl + +end point + +variables {W x₁ y₁} + +lemma equation_neg_iff : W.equation x₁ (W.neg_Y x₁ y₁) ↔ W.equation x₁ y₁ := +by { rw [equation_iff, equation_iff, neg_Y], congr' 2, ring1 } + +lemma equation_neg_of (h : W.equation x₁ $ W.neg_Y x₁ y₁) : W.equation x₁ y₁ := +equation_neg_iff.mp h + +/-- The negation of an affine point in `W` lies in `W`. -/ +lemma equation_neg (h : W.equation x₁ y₁) : W.equation x₁ $ W.neg_Y x₁ y₁ := equation_neg_iff.mpr h + +lemma nonsingular_neg_iff : W.nonsingular x₁ (W.neg_Y x₁ y₁) ↔ W.nonsingular x₁ y₁ := +begin + rw [nonsingular_iff, equation_neg_iff, ← neg_Y, neg_Y_neg_Y, ← @ne_comm _ y₁, nonsingular_iff], + exact and_congr_right' ((iff_congr not_and_distrib.symm not_and_distrib.symm).mpr $ + not_iff_not_of_iff $ and_congr_left $ λ h, by rw [← h]) +end + +lemma nonsingular_neg_of (h : W.nonsingular x₁ $ W.neg_Y x₁ y₁) : W.nonsingular x₁ y₁ := +nonsingular_neg_iff.mp h + +/-- The negation of a nonsingular affine point in `W` is nonsingular. -/ +lemma nonsingular_neg (h : W.nonsingular x₁ y₁) : W.nonsingular x₁ $ W.neg_Y x₁ y₁ := +nonsingular_neg_iff.mpr h + +namespace point + +/-- The negation of a nonsingular rational point. + +Given a nonsingular rational point `P`, use `-P` instead of `neg P`. -/ +def neg : W.point → W.point +| 0 := 0 +| (some h) := some $ nonsingular_neg h + +instance : has_neg W.point := ⟨neg⟩ + +@[simp] lemma neg_def (P : W.point) : P.neg = -P := rfl + +@[simp] lemma neg_zero : (-0 : W.point) = 0 := rfl + +@[simp] lemma neg_some (h : W.nonsingular x₁ y₁) : -some h = some (nonsingular_neg h) := rfl + +instance : has_involutive_neg W.point := ⟨neg, by { rintro (_ | _), { refl }, { simp, ring1 } }⟩ + +end point + +end basic + +section addition + +/-! ### Slopes of lines through nonsingular rational points on a Weierstrass curve -/ + +open_locale classical + +variables {F : Type u} [field F] (W : weierstrass_curve F) (x₁ x₂ y₁ y₂ : F) + +/-- The slope of the line through two affine points $(x_1, y_1)$ and $(x_2, y_2)$ in `W`. +If $x_1 \ne x_2$, then this line is the secant of `W` through $(x_1, y_1)$ and $(x_2, y_2)$, +and has slope $(y_1 - y_2) / (x_1 - x_2)$. Otherwise, if $y_1 \ne -y_1 - a_1x_1 - a_3$, +then this line is the tangent of `W` at $(x_1, y_1) = (x_2, y_2)$, and has slope +$(3x_1^2 + 2a_2x_1 + a_4 - a_1y_1) / (2y_1 + a_1x_1 + a_3)$. Otherwise, this line is vertical, +and has undefined slope, in which case this function returns the value 0. + +This depends on `W`, and has argument order: $x_1$, $x_2$, $y_1$, $y_2$. -/ +noncomputable def slope : F := +if hx : x₁ = x₂ then if hy : y₁ = W.neg_Y x₂ y₂ then 0 +else (3 * x₁ ^ 2 + 2 * W.a₂ * x₁ + W.a₄ - W.a₁ * y₁) / (y₁ - W.neg_Y x₁ y₁) +else (y₁ - y₂) / (x₁ - x₂) + +variables {W x₁ x₂ y₁ y₂} (h₁ : W.nonsingular x₁ y₁) (h₂ : W.nonsingular x₂ y₂) + (h₁' : W.equation x₁ y₁) (h₂' : W.equation x₂ y₂) + +@[simp] lemma slope_of_Y_ne (hx : x₁ = x₂) (hy : y₁ ≠ W.neg_Y x₂ y₂) : + W.slope x₁ x₂ y₁ y₂ = (3 * x₁ ^ 2 + 2 * W.a₂ * x₁ + W.a₄ - W.a₁ * y₁) / (y₁ - W.neg_Y x₁ y₁) := +by rw [slope, dif_pos hx, dif_neg hy] + +@[simp] lemma slope_of_X_ne (hx : x₁ ≠ x₂) : W.slope x₁ x₂ y₁ y₂ = (y₁ - y₂) / (x₁ - x₂) := +by rw [slope, dif_neg hx] + +lemma slope_of_Y_ne_eq_eval (hx : x₁ = x₂) (hy : y₁ ≠ W.neg_Y x₂ y₂) : + W.slope x₁ x₂ y₁ y₂ + = -(W.polynomial_X.eval $ C y₁).eval x₁ / (W.polynomial_Y.eval $ C y₁).eval x₁ := +by { rw [slope_of_Y_ne hx hy, eval_polynomial_X, neg_sub], congr' 1, rw [neg_Y, eval_polynomial_Y], + ring1 } + +include h₁' h₂' + +lemma Y_eq_of_X_eq (hx : x₁ = x₂) : y₁ = y₂ ∨ y₁ = W.neg_Y x₂ y₂ := +begin + rw [equation_iff] at h₁' h₂', + rw [← sub_eq_zero, ← @sub_eq_zero _ _ y₁, ← mul_eq_zero, neg_Y], + linear_combination h₁' - h₂' with { normalization_tactic := `[rw [hx], ring1] } +end + +lemma Y_eq_of_Y_ne (hx : x₁ = x₂) (hy : y₁ ≠ W.neg_Y x₂ y₂) : y₁ = y₂ := +or.resolve_right (Y_eq_of_X_eq h₁' h₂' hx) hy + +lemma add_polynomial_slope (hxy : x₁ = x₂ → y₁ ≠ W.neg_Y x₂ y₂) : + W.add_polynomial x₁ y₁ (W.slope x₁ x₂ y₁ y₂) + = -((X - C x₁) * (X - C x₂) * (X - C (W.add_X x₁ x₂ $ W.slope x₁ x₂ y₁ y₂))) := +begin + rw [add_polynomial_eq, neg_inj, cubic.prod_X_sub_C_eq, cubic.to_poly_injective], + by_cases hx : x₁ = x₂, + { rcases ⟨hx, Y_eq_of_Y_ne h₁' h₂' hx (hxy hx)⟩ with ⟨rfl, rfl⟩, + rw [equation_iff] at h₁' h₂', + rw [slope_of_Y_ne rfl $ hxy rfl], + rw [neg_Y, ← sub_ne_zero] at hxy, + ext, + { refl }, + { simp only [add_X], + ring1 }, + { field_simp [hxy rfl], + ring1 }, + { linear_combination -h₁' with { normalization_tactic := `[field_simp [hxy rfl], ring1] } } }, + { rw [equation_iff] at h₁' h₂', + rw [slope_of_X_ne hx], + rw [← sub_eq_zero] at hx, + ext, + { refl }, + { simp only [add_X], + ring1 }, + { apply mul_right_injective₀ hx, + linear_combination h₂' - h₁' with { normalization_tactic := `[field_simp [hx], ring1] } }, + { apply mul_right_injective₀ hx, + linear_combination x₂ * h₁' - x₁ * h₂' + with { normalization_tactic := `[field_simp [hx], ring1] } } } +end + +lemma derivative_add_polynomial_slope (hxy : x₁ = x₂ → y₁ ≠ W.neg_Y x₂ y₂) : + derivative (W.add_polynomial x₁ y₁ $ W.slope x₁ x₂ y₁ y₂) + = -((X - C x₁) * (X - C x₂) + (X - C x₁) * (X - C (W.add_X x₁ x₂ $ W.slope x₁ x₂ y₁ y₂)) + + (X - C x₂) * (X - C (W.add_X x₁ x₂ $ W.slope x₁ x₂ y₁ y₂))) := +by { rw [add_polynomial_slope h₁' h₂' hxy], derivative_simp, ring1 } + +/-! ### The addition law on nonsingular rational points on a Weierstrass curve -/ + +/-- The addition of two affine points in `W` on a sloped line, +before applying the final negation that maps $Y$ to $-Y - a_1X - a_3$, lies in `W`. -/ +lemma equation_add' (hxy : x₁ = x₂ → y₁ ≠ W.neg_Y x₂ y₂) : + W.equation (W.add_X x₁ x₂ $ W.slope x₁ x₂ y₁ y₂) (W.add_Y' x₁ x₂ y₁ $ W.slope x₁ x₂ y₁ y₂) := +by { rw [equation_add_iff, add_polynomial_slope h₁' h₂' hxy], eval_simp, + rw [neg_eq_zero, sub_self, mul_zero] } + +/-- The addition of two affine points in `W` on a sloped line lies in `W`. -/ +lemma equation_add (hxy : x₁ = x₂ → y₁ ≠ W.neg_Y x₂ y₂) : + W.equation (W.add_X x₁ x₂ $ W.slope x₁ x₂ y₁ y₂) (W.add_Y x₁ x₂ y₁ $ W.slope x₁ x₂ y₁ y₂) := +equation_neg $ equation_add' h₁' h₂' hxy + +omit h₁' h₂' + +include h₁ h₂ + +/-- The addition of two nonsingular affine points in `W` on a sloped line, +before applying the final negation that maps $Y$ to $-Y - a_1X - a_3$, is nonsingular. -/ +lemma nonsingular_add' (hxy : x₁ = x₂ → y₁ ≠ W.neg_Y x₂ y₂) : + W.nonsingular (W.add_X x₁ x₂ $ W.slope x₁ x₂ y₁ y₂) (W.add_Y' x₁ x₂ y₁ $ W.slope x₁ x₂ y₁ y₂) := +begin + by_cases hx₁ : W.add_X x₁ x₂ (W.slope x₁ x₂ y₁ y₂) = x₁, + { rwa [add_Y', hx₁, sub_self, mul_zero, zero_add] }, + { by_cases hx₂ : W.add_X x₁ x₂ (W.slope x₁ x₂ y₁ y₂) = x₂, + { by_cases hx : x₁ = x₂, + { subst hx, + contradiction }, + { rwa [add_Y', ← neg_sub, mul_neg, hx₂, slope_of_X_ne hx, + div_mul_cancel _ $ sub_ne_zero_of_ne hx, neg_sub, sub_add_cancel] } }, + { apply W.nonsingular_add_of_eval_derivative_ne_zero _ _ _ _ (equation_add' h₁.1 h₂.1 hxy), + rw [derivative_add_polynomial_slope h₁.left h₂.left hxy], + eval_simp, + simpa only [neg_ne_zero, sub_self, mul_zero, add_zero] + using mul_ne_zero (sub_ne_zero_of_ne hx₁) (sub_ne_zero_of_ne hx₂) } } +end + +/-- The addition of two nonsingular affine points in `W` on a sloped line is nonsingular. -/ +lemma nonsingular_add (hxy : x₁ = x₂ → y₁ ≠ W.neg_Y x₂ y₂) : + W.nonsingular (W.add_X x₁ x₂ $ W.slope x₁ x₂ y₁ y₂) (W.add_Y x₁ x₂ y₁ $ W.slope x₁ x₂ y₁ y₂) := +nonsingular_neg $ nonsingular_add' h₁ h₂ hxy + +omit h₁ h₂ + +namespace point + +/-- The addition of two nonsingular rational points. + +Given two nonsingular rational points `P` and `Q`, use `P + Q` instead of `add P Q`. -/ +noncomputable def add : W.point → W.point → W.point +| 0 P := P +| P 0 := P +| (@some _ _ _ x₁ y₁ h₁) (@some _ _ _ x₂ y₂ h₂) := +if hx : x₁ = x₂ then if hy : y₁ = W.neg_Y x₂ y₂ then 0 +else some $ nonsingular_add h₁ h₂ $ λ _, hy +else some $ nonsingular_add h₁ h₂ $ λ h, (hx h).elim + +noncomputable instance : has_add W.point := ⟨add⟩ + +@[simp] lemma add_def (P Q : W.point) : P.add Q = P + Q := rfl + +noncomputable instance : add_zero_class W.point := +⟨0, (+), by rintro (_ | _); refl, by rintro (_ | _); refl⟩ + +@[simp] lemma some_add_some_of_Y_eq (hx : x₁ = x₂) (hy : y₁ = W.neg_Y x₂ y₂) : + some h₁ + some h₂ = 0 := +by rw [← add_def, add, dif_pos hx, dif_pos hy] + +@[simp] lemma some_add_self_of_Y_eq (hy : y₁ = W.neg_Y x₁ y₁) : some h₁ + some h₁ = 0 := +some_add_some_of_Y_eq h₁ h₁ rfl hy + +@[simp] lemma some_add_some_of_Y_ne (hx : x₁ = x₂) (hy : y₁ ≠ W.neg_Y x₂ y₂) : + some h₁ + some h₂ = some (nonsingular_add h₁ h₂ $ λ _, hy) := +by rw [← add_def, add, dif_pos hx, dif_neg hy] + +lemma some_add_some_of_Y_ne' (hx : x₁ = x₂) (hy : y₁ ≠ W.neg_Y x₂ y₂) : + some h₁ + some h₂ = -some (nonsingular_add' h₁ h₂ $ λ _, hy) := +some_add_some_of_Y_ne h₁ h₂ hx hy + +@[simp] lemma some_add_self_of_Y_ne (hy : y₁ ≠ W.neg_Y x₁ y₁) : + some h₁ + some h₁ = some (nonsingular_add h₁ h₁ $ λ _, hy) := +some_add_some_of_Y_ne h₁ h₁ rfl hy + +lemma some_add_self_of_Y_ne' (hy : y₁ ≠ W.neg_Y x₁ y₁) : + some h₁ + some h₁ = -some (nonsingular_add' h₁ h₁ $ λ _, hy) := +some_add_some_of_Y_ne h₁ h₁ rfl hy + +@[simp] lemma some_add_some_of_X_ne (hx : x₁ ≠ x₂) : + some h₁ + some h₂ = some (nonsingular_add h₁ h₂ $ λ h, (hx h).elim) := +by rw [← add_def, add, dif_neg hx] + +lemma some_add_some_of_X_ne' (hx : x₁ ≠ x₂) : + some h₁ + some h₂ = -some (nonsingular_add' h₁ h₂ $ λ h, (hx h).elim) := +some_add_some_of_X_ne h₁ h₂ hx + +end point + +end addition + +section group + +/-! ### The axioms for nonsingular rational points on a Weierstrass curve -/ + +variables {F : Type u} [field F] {W : weierstrass_curve F} + +namespace point + +@[simp] lemma add_eq_zero (P Q : W.point) : P + Q = 0 ↔ P = -Q := +begin + rcases ⟨P, Q⟩ with ⟨_ | @⟨x₁, y₁, h₁⟩, _ | @⟨x₂, y₂, h₂⟩⟩, + any_goals { refl }, + { rw [zero_def, zero_add, eq_neg_iff_eq_neg, neg_zero] }, + { simp only [neg_some], + split, + { intro h, + by_cases hx : x₁ = x₂, + { by_cases hy : y₁ = W.neg_Y x₂ y₂, + { exact ⟨hx, hy⟩ }, + { rw [some_add_some_of_Y_ne h₁ h₂ hx hy] at h, + contradiction } }, + { rw [some_add_some_of_X_ne h₁ h₂ hx] at h, + contradiction } }, + { exact λ ⟨hx, hy⟩, some_add_some_of_Y_eq h₁ h₂ hx hy } } +end + +@[simp] lemma add_left_neg (P : W.point) : -P + P = 0 := by rw [add_eq_zero] + +@[simp] lemma neg_add_eq_zero (P Q : W.point) : -P + Q = 0 ↔ P = Q := by rw [add_eq_zero, neg_inj] + +end point + +end group + +end weierstrass_curve + +namespace elliptic_curve + +/-! ### Rational points on an elliptic curve -/ + +namespace point + +variables {R : Type} [nontrivial R] [comm_ring R] (E : elliptic_curve R) + +/-- An affine point on an elliptic curve `E` over `R`. -/ +def mk {x y : R} (h : E.equation x y) : E.point := weierstrass_curve.point.some $ E.nonsingular h + +end point + +end elliptic_curve diff --git a/src/algebraic_geometry/elliptic_curve/weierstrass.lean b/src/algebraic_geometry/elliptic_curve/weierstrass.lean index 704fd430cd765..b4e8cc54f81e1 100644 --- a/src/algebraic_geometry/elliptic_curve/weierstrass.lean +++ b/src/algebraic_geometry/elliptic_curve/weierstrass.lean @@ -39,7 +39,7 @@ splitting field of `R` are precisely the $X$-coordinates of the non-zero 2-torsi * `weierstrass_curve.nonsingular`: the nonsingular condition at a point on a Weierstrass curve. * `weierstrass_curve.coordinate_ring`: the coordinate ring of a Weierstrass curve. * `weierstrass_curve.function_field`: the function field of a Weierstrass curve. - * `weierstrass_curve.basis`: the power basis of the coordinate ring as an `R[X]`-algebra. + * `weierstrass_curve.coordinate_ring.basis`: the power basis of the coordinate ring over `R[X]`. * `elliptic_curve`: an elliptic curve over a commutative ring. * `elliptic_curve.j`: the j-invariant of an elliptic curve. @@ -51,8 +51,8 @@ splitting field of `R` are precisely the $X$-coordinates of the non-zero 2-torsi if its discriminant is non-zero. * `weierstrass_curve.coordinate_ring.is_domain`: the coordinate ring of a Weierstrass curve is an integral domain. - * `weierstrass_curve.degree_norm_smul_basis`: the degree of the norm of an element in the - coordinate ring as an `R[X]`-algebra in terms of the power basis. + * `weierstrass_curve.coordinate_ring.degree_norm_smul_basis`: the degree of the norm of an element + in the coordinate ring in terms of the power basis. * `elliptic_curve.nonsingular`: an elliptic curve is nonsingular at every point. * `elliptic_curve.variable_change_j`: the j-invariant of an elliptic curve is invariant under an admissible linear change of variables. @@ -240,9 +240,14 @@ lemma two_torsion_polynomial_disc_ne_zero [nontrivial R] [invertible (2 : R)] (h end torsion_polynomial +localized "notation (name := outer_variable) `Y` := polynomial.X" in polynomial_polynomial + +localized "notation (name := polynomial_polynomial) R`[X][Y]` := polynomial (polynomial R)" + in polynomial_polynomial + section polynomial -/-! ### Weierstrass polynomials and equations -/ +/-! ### Weierstrass equations -/ open polynomial @@ -250,9 +255,11 @@ open_locale polynomial /-- The polynomial $W(X, Y) := Y^2 + a_1XY + a_3Y - (X^3 + a_2X^2 + a_4X + a_6)$ associated to a Weierstrass curve `W` over `R`. For ease of polynomial manipulation, this is represented as a term -of type `R[X][X]`, where the inner variable represents $X$ and the outer variable represents $Y$. -/ -protected noncomputable def polynomial : R[X][X] := -X ^ 2 + C (C W.a₁ * X + C W.a₃) * X - C (X ^ 3 + C W.a₂ * X ^ 2 + C W.a₄ * X + C W.a₆) +of type `R[X][X]`, where the inner variable represents $X$ and the outer variable represents $Y$. +For clarity, the alternative notations `Y` and `R[X][Y]` are provided in the `polynomial_polynomial` +locale to represent the outer variable and the bivariate polynomial ring `R[X][X]` respectively. -/ +protected noncomputable def polynomial : R[X][Y] := +Y ^ 2 + C (C W.a₁ * X + C W.a₃) * Y - C (X ^ 3 + C W.a₂ * X ^ 2 + C W.a₄ * X + C W.a₆) lemma polynomial_eq : W.polynomial = cubic.to_poly ⟨0, 1, cubic.to_poly ⟨0, 0, W.a₁, W.a₃⟩, cubic.to_poly ⟨-1, -W.a₂, -W.a₄, -W.a₆⟩⟩ := @@ -285,15 +292,15 @@ begin end @[simp] lemma eval_polynomial (x y : R) : - eval x (eval (C y) W.polynomial) + (W.polynomial.eval $ C y).eval x = y ^ 2 + W.a₁ * x * y + W.a₃ * y - (x ^ 3 + W.a₂ * x ^ 2 + W.a₄ * x + W.a₆) := by { simp only [weierstrass_curve.polynomial], eval_simp, rw [add_mul, ← add_assoc] } -@[simp] lemma eval_polynomial_zero : eval 0 (eval 0 W.polynomial) = -W.a₆ := +@[simp] lemma eval_polynomial_zero : (W.polynomial.eval 0).eval 0 = -W.a₆ := by simp only [← C_0, eval_polynomial, zero_add, zero_sub, mul_zero, zero_pow (nat.zero_lt_succ _)] /-- The proposition that an affine point $(x, y)$ lies in `W`. In other words, $W(x, y) = 0$. -/ -def equation (x y : R) : Prop := eval x (eval (C y) W.polynomial) = 0 +def equation (x y : R) : Prop := (W.polynomial.eval $ C y).eval x = 0 lemma equation_iff' (x y : R) : W.equation x y ↔ y ^ 2 + W.a₁ * x * y + W.a₃ * y - (x ^ 3 + W.a₂ * x ^ 2 + W.a₄ * x + W.a₆) = 0 := @@ -316,51 +323,57 @@ end /-! ### Nonsingularity of Weierstrass curves -/ -/-- The partial derivative $W_X(X, Y)$ of $W(X, Y)$ with respect to $X$. -/ -noncomputable def polynomial_X : R[X][X] := -C (C W.a₁) * X - C (C 3 * X ^ 2 + C (2 * W.a₂) * X + C W.a₄) +/-- The partial derivative $W_X(X, Y)$ of $W(X, Y)$ with respect to $X$. + +TODO: define this in terms of `polynomial.derivative`. -/ +noncomputable def polynomial_X : R[X][Y] := +C (C W.a₁) * Y - C (C 3 * X ^ 2 + C (2 * W.a₂) * X + C W.a₄) @[simp] lemma eval_polynomial_X (x y : R) : - eval x (eval (C y) W.polynomial_X) = W.a₁ * y - (3 * x ^ 2 + 2 * W.a₂ * x + W.a₄) := + (W.polynomial_X.eval $ C y).eval x = W.a₁ * y - (3 * x ^ 2 + 2 * W.a₂ * x + W.a₄) := by { simp only [polynomial_X], eval_simp } -@[simp] lemma eval_polynomial_X_zero : eval 0 (eval 0 W.polynomial_X) = -W.a₄ := +@[simp] lemma eval_polynomial_X_zero : (W.polynomial_X.eval 0).eval 0 = -W.a₄ := by simp only [← C_0, eval_polynomial_X, zero_add, zero_sub, mul_zero, zero_pow zero_lt_two] -/-- The partial derivative $W_Y(X, Y)$ of $W(X, Y)$ with respect to $Y$. -/ -noncomputable def polynomial_Y : R[X][X] := C (C 2) * X + C (C W.a₁ * X + C W.a₃) +/-- The partial derivative $W_Y(X, Y)$ of $W(X, Y)$ with respect to $Y$. + +TODO: define this in terms of `polynomial.derivative`. -/ +noncomputable def polynomial_Y : R[X][Y] := C (C 2) * Y + C (C W.a₁ * X + C W.a₃) @[simp] lemma eval_polynomial_Y (x y : R) : - eval x (eval (C y) W.polynomial_Y) = 2 * y + W.a₁ * x + W.a₃ := + (W.polynomial_Y.eval $ C y).eval x = 2 * y + W.a₁ * x + W.a₃ := by { simp only [polynomial_Y], eval_simp, rw [← add_assoc] } -@[simp] lemma eval_polynomial_Y_zero : eval 0 (eval 0 W.polynomial_Y) = W.a₃ := +@[simp] lemma eval_polynomial_Y_zero : (W.polynomial_Y.eval 0).eval 0 = W.a₃ := by simp only [← C_0, eval_polynomial_Y, zero_add, mul_zero] /-- The proposition that an affine point $(x, y)$ on `W` is nonsingular. In other words, either $W_X(x, y) \ne 0$ or $W_Y(x, y) \ne 0$. -/ def nonsingular (x y : R) : Prop := -eval x (eval (C y) W.polynomial_X) ≠ 0 ∨ eval x (eval (C y) W.polynomial_Y) ≠ 0 +W.equation x y ∧ ((W.polynomial_X.eval $ C y).eval x ≠ 0 ∨ (W.polynomial_Y.eval $ C y).eval x ≠ 0) lemma nonsingular_iff' (x y : R) : W.nonsingular x y - ↔ W.a₁ * y - (3 * x ^ 2 + 2 * W.a₂ * x + W.a₄) ≠ 0 ∨ 2 * y + W.a₁ * x + W.a₃ ≠ 0 := -by rw [nonsingular, eval_polynomial_X, eval_polynomial_Y] + ↔ W.equation x y + ∧ (W.a₁ * y - (3 * x ^ 2 + 2 * W.a₂ * x + W.a₄) ≠ 0 ∨ 2 * y + W.a₁ * x + W.a₃ ≠ 0) := +by rw [nonsingular, equation_iff', eval_polynomial_X, eval_polynomial_Y] @[simp] lemma nonsingular_iff (x y : R) : - W.nonsingular x y ↔ W.a₁ * y ≠ 3 * x ^ 2 + 2 * W.a₂ * x + W.a₄ ∨ y ≠ -y - W.a₁ * x - W.a₃ := -by { rw [nonsingular_iff', sub_ne_zero, ← @sub_ne_zero _ _ y], congr' 3; ring1 } + W.nonsingular x y + ↔ W.equation x y ∧ (W.a₁ * y ≠ 3 * x ^ 2 + 2 * W.a₂ * x + W.a₄ ∨ y ≠ -y - W.a₁ * x - W.a₃) := +by { rw [nonsingular_iff', sub_ne_zero, ← @sub_ne_zero _ _ y], congr' 4; ring1 } -@[simp] lemma nonsingular_zero : W.nonsingular 0 0 ↔ W.a₃ ≠ 0 ∨ W.a₄ ≠ 0 := -by rw [nonsingular, C_0, eval_polynomial_X_zero, neg_ne_zero, eval_polynomial_Y_zero, or_comm] +@[simp] lemma nonsingular_zero : W.nonsingular 0 0 ↔ W.a₆ = 0 ∧ (W.a₃ ≠ 0 ∨ W.a₄ ≠ 0) := +by rw [nonsingular, equation_zero, C_0, eval_polynomial_X_zero, neg_ne_zero, eval_polynomial_Y_zero, + or_comm] lemma nonsingular_iff_variable_change (x y : R) : W.nonsingular x y ↔ (W.variable_change 1 x 0 y).nonsingular 0 0 := begin - rw [nonsingular_iff', ← neg_ne_zero, or_comm, nonsingular_zero, variable_change_a₃, - variable_change_a₄, inv_one, units.coe_one], - congr' 3, - all_goals { ring1 } + rw [nonsingular_iff', equation_iff_variable_change, equation_zero, ← neg_ne_zero, or_comm, + nonsingular_zero, variable_change_a₃, variable_change_a₄, inv_one, units.coe_one], + congr' 4; ring1 end lemma nonsingular_zero_of_Δ_ne_zero (h : W.equation 0 0) (hΔ : W.Δ ≠ 0) : W.nonsingular 0 0 := @@ -388,7 +401,7 @@ https://leanprover.zulipchat.com/#narrow/stream/116395-maths/topic/.E2.9C.94.20c @[derive [inhabited, comm_ring]] def coordinate_ring : Type u := adjoin_root W.polynomial /-- The function field $R(W) := \mathrm{Frac}(R[W])$ of `W`. -/ -@[reducible] def function_field : Type u := fraction_ring W.coordinate_ring +abbreviation function_field : Type u := fraction_ring W.coordinate_ring namespace coordinate_ring @@ -411,7 +424,7 @@ adjoin_root.mk_ne_zero_of_nat_degree_lt W.monic_polynomial (C_ne_zero.mpr $ X_su by { rw [nat_degree_polynomial, nat_degree_C], norm_num1 } /-- The class of the element $Y - y(X)$ in $R[W]$ for some $y(X) \in R[X]$. -/ -@[simp] noncomputable def Y_class : W.coordinate_ring := adjoin_root.mk W.polynomial $ X - C y +@[simp] noncomputable def Y_class : W.coordinate_ring := adjoin_root.mk W.polynomial $ Y - C y lemma Y_class_ne_zero [nontrivial R] : Y_class W y ≠ 0 := adjoin_root.mk_ne_zero_of_nat_degree_lt W.monic_polynomial (X_sub_C_ne_zero y) $ @@ -433,16 +446,13 @@ instance : is_scalar_tower R R[X] W.coordinate_ring := ideal.quotient.is_scalar_ instance [subsingleton R] : subsingleton W.coordinate_ring := module.subsingleton R[X] _ -section -open_locale classical /-- The basis $\{1, Y\}$ for the coordinate ring $R[W]$ over the polynomial ring $R[X]$. Given a Weierstrass curve `W`, write `W^.coordinate_ring.basis` for this basis. -/ protected noncomputable def basis : basis (fin 2) R[X] W.coordinate_ring := (subsingleton_or_nontrivial R).by_cases (λ _, by exactI default) $ λ _, by exactI - (basis.reindex (adjoin_root.power_basis' W.monic_polynomial).basis $ - fin_congr $ W.nat_degree_polynomial) -end + ((adjoin_root.power_basis' W.monic_polynomial).basis.reindex $ + fin_congr W.nat_degree_polynomial) lemma basis_apply (n : fin 2) : W^.coordinate_ring.basis n = (adjoin_root.power_basis' W.monic_polynomial).gen ^ (n : ℕ) := @@ -450,16 +460,16 @@ begin classical, nontriviality R, simpa only [coordinate_ring.basis, or.by_cases, dif_neg (not_subsingleton R), - basis.reindex_apply, power_basis.basis_eq_pow] + basis.reindex_apply, power_basis.basis_eq_pow] end lemma basis_zero : W^.coordinate_ring.basis 0 = 1 := by simpa only [basis_apply] using pow_zero _ -lemma basis_one : W^.coordinate_ring.basis 1 = adjoin_root.mk W.polynomial X := +lemma basis_one : W^.coordinate_ring.basis 1 = adjoin_root.mk W.polynomial Y := by simpa only [basis_apply] using pow_one _ @[simp] lemma coe_basis : - (W^.coordinate_ring.basis : fin 2 → W.coordinate_ring) = ![1, adjoin_root.mk W.polynomial X] := + (W^.coordinate_ring.basis : fin 2 → W.coordinate_ring) = ![1, adjoin_root.mk W.polynomial Y] := by { ext n, fin_cases n, exacts [basis_zero W, basis_one W] } variable {W} @@ -468,7 +478,7 @@ lemma smul (x : R[X]) (y : W.coordinate_ring) : x • y = adjoin_root.mk W.polyn (algebra_map_smul W.coordinate_ring x y).symm lemma smul_basis_eq_zero {p q : R[X]} - (hpq : p • 1 + q • adjoin_root.mk W.polynomial X = 0) : p = 0 ∧ q = 0 := + (hpq : p • 1 + q • adjoin_root.mk W.polynomial Y = 0) : p = 0 ∧ q = 0 := begin have h := fintype.linear_independent_iff.mp (coordinate_ring.basis W).linear_independent ![p, q], erw [fin.sum_univ_succ, basis_zero, fin.sum_univ_one, basis_one] at h, @@ -476,7 +486,7 @@ begin end lemma exists_smul_basis_eq (x : W.coordinate_ring) : - ∃ p q : R[X], p • 1 + q • adjoin_root.mk W.polynomial X = x := + ∃ p q : R[X], p • 1 + q • adjoin_root.mk W.polynomial Y = x := begin have h := (coordinate_ring.basis W).sum_equiv_fun x, erw [fin.sum_univ_succ, fin.sum_univ_one, basis_zero, basis_one] at h, @@ -486,17 +496,17 @@ end variable (W) lemma smul_basis_mul_C (p q : R[X]) : - (p • 1 + q • adjoin_root.mk W.polynomial X) * adjoin_root.mk W.polynomial (C y) - = ((p * y) • 1 + (q * y) • adjoin_root.mk W.polynomial X) := + (p • 1 + q • adjoin_root.mk W.polynomial Y) * adjoin_root.mk W.polynomial (C y) + = ((p * y) • 1 + (q * y) • adjoin_root.mk W.polynomial Y) := by { simp only [smul, map_mul], ring1 } lemma smul_basis_mul_Y (p q : R[X]) : - (p • 1 + q • adjoin_root.mk W.polynomial X) * adjoin_root.mk W.polynomial X + (p • 1 + q • adjoin_root.mk W.polynomial Y) * adjoin_root.mk W.polynomial Y = (q * (X ^ 3 + C W.a₂ * X ^ 2 + C W.a₄ * X + C W.a₆)) • 1 - + (p - q * (C W.a₁ * X + C W.a₃)) • adjoin_root.mk W.polynomial X := + + (p - q * (C W.a₁ * X + C W.a₃)) • adjoin_root.mk W.polynomial Y := begin - have Y_sq : adjoin_root.mk W.polynomial X ^ 2 = adjoin_root.mk W.polynomial - (C (X ^ 3 + C W.a₂ * X ^ 2 + C W.a₄ * X + C W.a₆) - C (C W.a₁ * X + C W.a₃) * X) := + have Y_sq : adjoin_root.mk W.polynomial Y ^ 2 = adjoin_root.mk W.polynomial + (C (X ^ 3 + C W.a₂ * X ^ 2 + C W.a₄ * X + C W.a₆) - C (C W.a₁ * X + C W.a₃) * Y) := adjoin_root.mk_eq_mk.mpr ⟨1, by { simp only [weierstrass_curve.polynomial], ring1 }⟩, simp only [smul, add_mul, mul_assoc, ← sq, Y_sq, map_sub, map_mul], ring1 @@ -505,7 +515,7 @@ end /-! ### Norms on the coordinate ring -/ lemma norm_smul_basis (p q : R[X]) : - algebra.norm R[X] (p • 1 + q • adjoin_root.mk W.polynomial X) + algebra.norm R[X] (p • 1 + q • adjoin_root.mk W.polynomial Y) = p ^ 2 - p * q * (C W.a₁ * X + C W.a₃) - q ^ 2 * (X ^ 3 + C W.a₂ * X ^ 2 + C W.a₄ * X + C W.a₆) := begin @@ -517,14 +527,14 @@ begin end lemma coe_norm_smul_basis (p q : R[X]) : - ↑(algebra.norm R[X] $ p • 1 + q • adjoin_root.mk W.polynomial X) + ↑(algebra.norm R[X] $ p • 1 + q • adjoin_root.mk W.polynomial Y) = adjoin_root.mk W.polynomial ((C p + C q * X) * (C p + C q * (-X - C (C W.a₁ * X + C W.a₃)))) := adjoin_root.mk_eq_mk.mpr ⟨C q ^ 2, by { rw [norm_smul_basis, weierstrass_curve.polynomial], C_simp, ring1 }⟩ lemma degree_norm_smul_basis [is_domain R] (p q : R[X]) : - (algebra.norm R[X] $ p • 1 + q • adjoin_root.mk W.polynomial X).degree + (algebra.norm R[X] $ p • 1 + q • adjoin_root.mk W.polynomial Y).degree = max (2 • p.degree) (2 • q.degree + 3) := begin have hdp : (p ^ 2).degree = 2 • p.degree := degree_pow p 2, From dc6c365e751e34d100e80fe6e314c3c3e0fd2988 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 10 Feb 2023 18:00:48 +0000 Subject: [PATCH 0461/1291] chore(algebra/group_with_zero/units/lemmas): `div_div_cancel_left'` (#18421) This is the `group_with_zero` version of `div_div_cancel_left`. Forward-ported in https://github.com/leanprover-community/mathlib4/pull/2191 --- src/algebra/group_with_zero/units/lemmas.lean | 2 ++ src/algebra/hom/units.lean | 4 ++++ 2 files changed, 6 insertions(+) diff --git a/src/algebra/group_with_zero/units/lemmas.lean b/src/algebra/group_with_zero/units/lemmas.lean index 68ced6b72bfc9..2f0ecc5d24a70 100644 --- a/src/algebra/group_with_zero/units/lemmas.lean +++ b/src/algebra/group_with_zero/units/lemmas.lean @@ -113,6 +113,8 @@ hb.is_unit.div_eq_div_iff hd.is_unit lemma div_div_cancel' (ha : a ≠ 0) : a / (a / b) = b := ha.is_unit.div_div_cancel +lemma div_div_cancel_left' (ha : a ≠ 0) : a / b / a = b⁻¹ := ha.is_unit.div_div_cancel_left + lemma div_helper (b : G₀) (h : a ≠ 0) : 1 / (a * b) * a = 1 / b := by rw [div_mul_eq_mul_div, one_mul, div_mul_right _ h] diff --git a/src/algebra/hom/units.lean b/src/algebra/hom/units.lean index 71f0f7f067f89..00aff194746f0 100644 --- a/src/algebra/hom/units.lean +++ b/src/algebra/hom/units.lean @@ -318,5 +318,9 @@ by rw [←(hb.mul hd).mul_left_inj, ←mul_assoc, hb.div_mul_cancel, ←mul_asso @[to_additive] protected lemma div_div_cancel (h : is_unit a) : a / (a / b) = b := by rw [div_div_eq_mul_div, h.mul_div_cancel_left] +@[to_additive] protected lemma div_div_cancel_left (h : is_unit a) : + a / b / a = b⁻¹ := +by rw [div_eq_mul_inv, div_eq_mul_inv, mul_right_comm, h.mul_inv_cancel, one_mul] + end division_comm_monoid end is_unit From 806c0bb86f6128cfa2f702285727518eb5244390 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 10 Feb 2023 21:01:32 +0000 Subject: [PATCH 0462/1291] feat(analysis/normed_space/triv_sq_zero_ext): exponential of dual numbers (#18200) In order for convergence to be well-defined, we put the product topology on `triv_zero_sq_ext R M` via its obvious internal representation as a product. [Zulip thread](https://leanprover.zulipchat.com/#narrow/stream/116395-maths/topic/Exponentials.20in.20seminormed.20algebras/near/321870256). --- src/analysis/normed_space/dual_number.lean | 34 ++++ .../normed_space/triv_sq_zero_ext.lean | 116 ++++++++++++++ src/topology/instances/triv_sq_zero_ext.lean | 145 ++++++++++++++++++ 3 files changed, 295 insertions(+) create mode 100644 src/analysis/normed_space/dual_number.lean create mode 100644 src/analysis/normed_space/triv_sq_zero_ext.lean create mode 100644 src/topology/instances/triv_sq_zero_ext.lean diff --git a/src/analysis/normed_space/dual_number.lean b/src/analysis/normed_space/dual_number.lean new file mode 100644 index 0000000000000..fbe680a7fcad0 --- /dev/null +++ b/src/analysis/normed_space/dual_number.lean @@ -0,0 +1,34 @@ +/- +Copyright (c) 2023 Eric Wieser. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Eric Wieser +-/ +import algebra.dual_number +import analysis.normed_space.triv_sq_zero_ext + +/-! +# Results on `dual_number R` related to the norm + +These are just restatements of similar statements about `triv_sq_zero_ext R M`. + +## Main results + +* `exp_eps` + +-/ + +namespace dual_number +open triv_sq_zero_ext + +variables (𝕜 : Type*) {R : Type*} + +variables [is_R_or_C 𝕜] [normed_comm_ring R] [normed_algebra 𝕜 R] +variables [topological_ring R] [complete_space R] [t2_space R] + +@[simp] lemma exp_eps : exp 𝕜 (eps : dual_number R) = 1 + eps := +exp_inr _ _ + +@[simp] lemma exp_smul_eps (r : R) : exp 𝕜 (r • eps : dual_number R) = 1 + r • eps := +by rw [eps, ←inr_smul, exp_inr] + +end dual_number diff --git a/src/analysis/normed_space/triv_sq_zero_ext.lean b/src/analysis/normed_space/triv_sq_zero_ext.lean new file mode 100644 index 0000000000000..cf49d3b26b5a0 --- /dev/null +++ b/src/analysis/normed_space/triv_sq_zero_ext.lean @@ -0,0 +1,116 @@ +/- +Copyright (c) 2023 Eric Wieser. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Eric Wieser +-/ +import analysis.normed_space.basic +import analysis.normed_space.exponential +import topology.instances.triv_sq_zero_ext + +/-! +# Results on `triv_sq_zero_ext R M` related to the norm + +For now, this file contains results about `exp` for this type. + +TODO: actually define a sensible norm on `triv_sq_zero_ext R M`, so that we have access to lemmas +like `exp_add`. + +## Main results + +* `triv_sq_zero_ext.fst_exp` +* `triv_sq_zero_ext.snd_exp` +* `triv_sq_zero_ext.exp_inl` +* `triv_sq_zero_ext.exp_inr` + +-/ + +variables (𝕜 : Type*) {R M : Type*} + +local notation `tsze` := triv_sq_zero_ext + +namespace triv_sq_zero_ext + +section topology +variables [topological_space R] [topological_space M] + + +/-- If `exp R x.fst` converges to `e` then `exp R x` converges to `inl e + inr (e • x.snd)`. -/ +lemma has_sum_exp_series [field 𝕜] [char_zero 𝕜] [comm_ring R] + [add_comm_group M] [algebra 𝕜 R] [module R M] [module 𝕜 M] [is_scalar_tower 𝕜 R M] + [topological_ring R] [topological_add_group M] [has_continuous_smul R M] + (x : tsze R M) {e : R} (h : has_sum (λ n, exp_series 𝕜 R n (λ _, x.fst)) e) : + has_sum (λ n, exp_series 𝕜 (tsze R M) n (λ _, x)) (inl e + inr (e • x.snd)) := +begin + simp_rw [exp_series_apply_eq] at *, + conv + { congr, + funext, + rw [←inl_fst_add_inr_snd_eq (x ^ _), fst_pow, snd_pow, smul_add, ←inr_smul, + ←inl_smul, nsmul_eq_smul_cast 𝕜 n, smul_smul, inv_mul_eq_div, ←inv_div, ←smul_assoc], }, + refine (has_sum_inl M h).add (has_sum_inr M _), + apply has_sum.smul_const, + rw [←has_sum_nat_add_iff' 1], swap, apply_instance, + rw [finset.range_one, finset.sum_singleton, nat.cast_zero, div_zero, inv_zero, zero_smul, + sub_zero], + simp_rw [←nat.succ_eq_add_one, nat.pred_succ, nat.factorial_succ, nat.cast_mul, + ←nat.succ_eq_add_one, + mul_div_cancel_left _ ((@nat.cast_ne_zero 𝕜 _ _ _).mpr $ nat.succ_ne_zero _)], + exact h, +end + +end topology + +section normed_ring +variables [is_R_or_C 𝕜] [normed_comm_ring R] [add_comm_group M] +variables [normed_algebra 𝕜 R] [module R M] [module 𝕜 M] [is_scalar_tower 𝕜 R M] +variables [topological_space M] [topological_ring R] +variables [topological_add_group M] [has_continuous_smul R M] +variables [complete_space R] [t2_space R] [t2_space M] + +lemma exp_def (x : tsze R M) : exp 𝕜 x = inl (exp 𝕜 x.fst) + inr (exp 𝕜 x.fst • x.snd) := +begin + simp_rw [exp, formal_multilinear_series.sum], + refine (has_sum_exp_series 𝕜 x _).tsum_eq, + exact exp_series_has_sum_exp _, +end + +@[simp] lemma fst_exp (x : tsze R M) : fst (exp 𝕜 x) = exp 𝕜 x.fst := +by rw [exp_def, fst_add, fst_inl, fst_inr, add_zero] + +@[simp] lemma snd_exp (x : tsze R M) : snd (exp 𝕜 x) = exp 𝕜 x.fst • x.snd := +by rw [exp_def, snd_add, snd_inl, snd_inr, zero_add] + +@[simp] lemma exp_inl (x : R) : exp 𝕜 (inl x : tsze R M) = inl (exp 𝕜 x) := +by rw [exp_def, fst_inl, snd_inl, smul_zero, inr_zero, add_zero] + +@[simp] lemma exp_inr (m : M) : exp 𝕜 (inr m : tsze R M) = 1 + inr m := +by rw [exp_def, fst_inr, exp_zero, snd_inr, one_smul, inl_one] + +/-- Polar form of trivial-square-zero extension. -/ +lemma eq_smul_exp_of_invertible (x : tsze R M) [invertible x.fst] : + x = x.fst • exp 𝕜 (⅟x.fst • inr x.snd) := +by rw [←inr_smul, exp_inr, smul_add, ←inl_one, ←inl_smul, ←inr_smul, smul_eq_mul, mul_one, + smul_smul, mul_inv_of_self, one_smul, inl_fst_add_inr_snd_eq] + +end normed_ring + + +section normed_field +variables [is_R_or_C 𝕜] [normed_field R] [add_comm_group M] +variables [normed_algebra 𝕜 R] [module R M] [module 𝕜 M] [is_scalar_tower 𝕜 R M] +variables [topological_space M] [topological_ring R] +variables [topological_add_group M] [has_continuous_smul R M] +variables [complete_space R] [t2_space R] [t2_space M] + +/-- More convenient version of `triv_sq_zero_ext.eq_smul_exp_of_invertible` for when `R` is a +field. -/ +lemma eq_smul_exp_of_ne_zero (x : tsze R M) (hx : x.fst ≠ 0) : + x = x.fst • exp 𝕜 (x.fst⁻¹ • inr x.snd) := +begin + letI : invertible x.fst := invertible_of_nonzero hx, + exact eq_smul_exp_of_invertible _ _ +end + +end normed_field + +end triv_sq_zero_ext diff --git a/src/topology/instances/triv_sq_zero_ext.lean b/src/topology/instances/triv_sq_zero_ext.lean new file mode 100644 index 0000000000000..2ea5c325c5b64 --- /dev/null +++ b/src/topology/instances/triv_sq_zero_ext.lean @@ -0,0 +1,145 @@ +/- +Copyright (c) 2023 Eric Wieser. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Eric Wieser +-/ +import algebra.triv_sq_zero_ext +import topology.algebra.infinite_sum +import topology.algebra.module.basic + +/-! +# Topology on `triv_sq_zero_ext R M` + +The type `triv_sq_zero_ext R M` inherits the topology from `R × M`. + +Note that this is not the topology induced by the seminorm on the dual numbers suggested by +[this Math.SE answer](https://math.stackexchange.com/a/1056378/1896), which instead induces +the topology pulled back through the projection map `triv_sq_zero_ext.fst : tsze R M → R`. +Obviously, that topology is not Hausdorff and using it would result in `exp` converging to more than +one value. + +## Main results + +* `triv_sq_zero_ext.topological_ring`: the ring operations are continuous + +-/ + +variables {α S R M : Type*} + +local notation `tsze` := triv_sq_zero_ext + +namespace triv_sq_zero_ext +variables [topological_space R] [topological_space M] + +instance : topological_space (tsze R M) := +topological_space.induced fst ‹_› ⊓ topological_space.induced snd ‹_› + +instance [t2_space R] [t2_space M] : t2_space (tsze R M) := +prod.t2_space + +lemma nhds_def (x : tsze R M) : nhds x = (nhds x.fst).prod (nhds x.snd) := +by cases x; exact nhds_prod_eq +lemma nhds_inl [has_zero M] (x : R) : nhds (inl x : tsze R M) = (nhds x).prod (nhds 0) := nhds_def _ +lemma nhds_inr [has_zero R] (m : M) : nhds (inr m : tsze R M) = (nhds 0).prod (nhds m) := nhds_def _ + +lemma continuous_fst : continuous (fst : tsze R M → R) := continuous_fst +lemma continuous_snd : continuous (snd : tsze R M → M) := continuous_snd + +lemma continuous_inl [has_zero M] : continuous (inl : R → tsze R M) := +continuous_id.prod_mk continuous_const +lemma continuous_inr [has_zero R] : continuous (inr : M → tsze R M) := +continuous_const.prod_mk continuous_id + +lemma embedding_inl [has_zero M] : embedding (inl : R → tsze R M) := +embedding_of_embedding_compose continuous_inl continuous_fst embedding_id +lemma embedding_inr [has_zero R] : embedding (inr : M → tsze R M) := +embedding_of_embedding_compose continuous_inr continuous_snd embedding_id + +variables (R M) + +/-- `triv_sq_zero_ext.fst` as a continuous linear map. -/ +@[simps] +def fst_clm [comm_semiring R] [add_comm_monoid M] [module R M] : tsze R M →L[R] R := +{ to_fun := fst, + .. continuous_linear_map.fst R R M } + +/-- `triv_sq_zero_ext.snd` as a continuous linear map. -/ +@[simps] +def snd_clm [comm_semiring R] [add_comm_monoid M] [module R M] : tsze R M →L[R] M := +{ to_fun := snd, + cont := continuous_snd, + .. continuous_linear_map.snd R R M } + +/-- `triv_sq_zero_ext.inl` as a continuous linear map. -/ +@[simps] +def inl_clm [comm_semiring R] [add_comm_monoid M] [module R M] : R →L[R] tsze R M := +{ to_fun := inl, + .. continuous_linear_map.inl R R M } + +/-- `triv_sq_zero_ext.inr` as a continuous linear map. -/ +@[simps] +def inr_clm [comm_semiring R] [add_comm_monoid M] [module R M] : M →L[R] tsze R M := +{ to_fun := inr, + .. continuous_linear_map.inr R R M } + +variables {R M} + +instance [has_add R] [has_add M] + [has_continuous_add R] [has_continuous_add M] : + has_continuous_add (tsze R M) := +prod.has_continuous_add + +instance [has_mul R] [has_add M] [has_smul R M] + [has_continuous_mul R] [has_continuous_smul R M] [has_continuous_add M] : + has_continuous_mul (tsze R M) := +⟨((continuous_fst.comp _root_.continuous_fst).mul (continuous_fst.comp _root_.continuous_snd)) + .prod_mk $ + ((continuous_fst.comp _root_.continuous_fst).smul + (continuous_snd.comp _root_.continuous_snd)).add + ((continuous_fst.comp _root_.continuous_snd).smul + (continuous_snd.comp _root_.continuous_fst))⟩ + +instance [has_neg R] [has_neg M] + [has_continuous_neg R] [has_continuous_neg M] : + has_continuous_neg (tsze R M) := +prod.has_continuous_neg + +instance [semiring R] [add_comm_monoid M] [module R M] + [topological_semiring R] [has_continuous_add M] [has_continuous_smul R M] : + topological_semiring (tsze R M) := +{} + +instance [comm_ring R] [add_comm_group M] [module R M] + [topological_ring R] [topological_add_group M] [has_continuous_smul R M] : + topological_ring (tsze R M) := +{} + +instance [has_smul S R] [has_smul S M] + [has_continuous_const_smul S R] [has_continuous_const_smul S M] : + has_continuous_const_smul S (tsze R M) := +prod.has_continuous_const_smul + +instance [topological_space S] [has_smul S R] [has_smul S M] + [has_continuous_smul S R] [has_continuous_smul S M] : + has_continuous_smul S (tsze R M) := +prod.has_continuous_smul + +variables (M) + +lemma has_sum_inl [add_comm_monoid R] [add_comm_monoid M] {f : α → R} {a : R} (h : has_sum f a) : + has_sum (λ x, inl (f x)) (inl a : tsze R M) := +h.map (⟨inl, inl_zero _, inl_add _⟩ : R →+ tsze R M) continuous_inl + +lemma has_sum_inr [add_comm_monoid R] [add_comm_monoid M] {f : α → M} {a : M} (h : has_sum f a) : + has_sum (λ x, inr (f x)) (inr a : tsze R M) := +h.map (⟨inr, inr_zero _, inr_add _⟩ : M →+ tsze R M) continuous_inr + +lemma has_sum_fst [add_comm_monoid R] [add_comm_monoid M] {f : α → tsze R M} {a : tsze R M} + (h : has_sum f a) : has_sum (λ x, fst (f x)) (fst a) := +h.map (⟨fst, fst_zero, fst_add⟩ : tsze R M →+ R) continuous_fst + +lemma has_sum_snd [add_comm_monoid R] [add_comm_monoid M] {f : α → tsze R M} {a : tsze R M} + (h : has_sum f a) : has_sum (λ x, snd (f x)) (snd a) := +h.map (⟨snd, snd_zero, snd_add⟩ : tsze R M →+ M) continuous_snd + +end triv_sq_zero_ext From ad0089aca372256fe53dde13ca0dfea569bf5ac7 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Sat, 11 Feb 2023 02:09:58 +0000 Subject: [PATCH 0463/1291] chore(*): add mathlib4 synchronization comments (#18416) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.order.upper_lower` * `data.pi.interval` * `number_theory.class_number.admissible_absolute_value` * `topology.algebra.order.filter` * `topology.fiber_bundle.is_homeomorphic_trivial_bundle` * `topology.hom.open` --- src/algebra/order/upper_lower.lean | 3 +++ src/data/pi/interval.lean | 3 +++ src/number_theory/class_number/admissible_absolute_value.lean | 3 +++ src/topology/algebra/order/filter.lean | 3 +++ src/topology/fiber_bundle/is_homeomorphic_trivial_bundle.lean | 3 +++ src/topology/hom/open.lean | 3 +++ 6 files changed, 18 insertions(+) diff --git a/src/algebra/order/upper_lower.lean b/src/algebra/order/upper_lower.lean index 7bfeea35e4f51..3f36d91f00fa1 100644 --- a/src/algebra/order/upper_lower.lean +++ b/src/algebra/order/upper_lower.lean @@ -10,6 +10,9 @@ import order.upper_lower.basic /-! # Algebraic operations on upper/lower sets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Upper/lower sets are preserved under pointwise algebraic operations in ordered groups. -/ diff --git a/src/data/pi/interval.lean b/src/data/pi/interval.lean index 7388b3bbb69d2..c77e21f3ce2bd 100644 --- a/src/data/pi/interval.lean +++ b/src/data/pi/interval.lean @@ -9,6 +9,9 @@ import data.fintype.big_operators /-! # Intervals in a pi type +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file shows that (dependent) functions to locally finite orders equipped with the pointwise order are locally finite and calculates the cardinality of their intervals. -/ diff --git a/src/number_theory/class_number/admissible_absolute_value.lean b/src/number_theory/class_number/admissible_absolute_value.lean index 8324f2ed74cab..745d6a7b62df4 100644 --- a/src/number_theory/class_number/admissible_absolute_value.lean +++ b/src/number_theory/class_number/admissible_absolute_value.lean @@ -9,6 +9,9 @@ import algebra.order.euclidean_absolute_value /-! # Admissible absolute values + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. This file defines a structure `absolute_value.is_admissible` which we use to show the class number of the ring of integers of a global field is finite. diff --git a/src/topology/algebra/order/filter.lean b/src/topology/algebra/order/filter.lean index 31ced314235d9..92213dd2b724e 100644 --- a/src/topology/algebra/order/filter.lean +++ b/src/topology/algebra/order/filter.lean @@ -9,6 +9,9 @@ import topology.filter /-! # Topology on filters of a space with order topology +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove that `𝓝 (f x)` tends to `𝓝 filter.at_top` provided that `f` tends to `filter.at_top`, and similarly for `filter.at_bot`. -/ diff --git a/src/topology/fiber_bundle/is_homeomorphic_trivial_bundle.lean b/src/topology/fiber_bundle/is_homeomorphic_trivial_bundle.lean index bd2502ee9930d..6df4fa1f910d2 100644 --- a/src/topology/fiber_bundle/is_homeomorphic_trivial_bundle.lean +++ b/src/topology/fiber_bundle/is_homeomorphic_trivial_bundle.lean @@ -8,6 +8,9 @@ import topology.homeomorph /-! # Maps equivariantly-homeomorphic to projection in a product +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains the definition `is_homeomorphic_trivial_fiber_bundle F p`, a Prop saying that a map `p : Z → B` between topological spaces is a "trivial fiber bundle" in the sense that there exists a homeomorphism `h : Z ≃ₜ B × F` such that `proj x = (h x).1`. This is an abstraction which diff --git a/src/topology/hom/open.lean b/src/topology/hom/open.lean index 84aeb9a84c3d3..9225e3976caf1 100644 --- a/src/topology/hom/open.lean +++ b/src/topology/hom/open.lean @@ -8,6 +8,9 @@ import topology.continuous_function.basic /-! # Continuous open maps +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines bundled continuous open maps. We use the `fun_like` design, so each type of morphisms has a companion typeclass which is meant to From ddbfecf55a52098524dce4a061d1eadba9790f3c Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Sat, 11 Feb 2023 09:22:54 +0000 Subject: [PATCH 0464/1291] feat(topology/continuous_function): lemmas on infinite sums of continuous functions (#18424) A lemma (in three variants) about evaluating infinite sums of continuous functions at a point. Split off from #18392. --- src/topology/continuous_function/algebra.lean | 88 +++++++++++-------- 1 file changed, 51 insertions(+), 37 deletions(-) diff --git a/src/topology/continuous_function/algebra.lean b/src/topology/continuous_function/algebra.lean index d0f9763b02ed5..224f10c45667f 100644 --- a/src/topology/continuous_function/algebra.lean +++ b/src/topology/continuous_function/algebra.lean @@ -8,6 +8,7 @@ import topology.continuous_function.ordered import topology.algebra.uniform_group import topology.uniform_space.compact_convergence import topology.algebra.star +import topology.algebra.infinite_sum import algebra.algebra.pi import algebra.algebra.subalgebra.basic import tactic.field_simp @@ -180,51 +181,43 @@ end subtype namespace continuous_map +variables {α β : Type*} [topological_space α] [topological_space β] + @[to_additive] -instance {α : Type*} {β : Type*} [topological_space α] - [topological_space β] [semigroup β] [has_continuous_mul β] : semigroup C(α, β) := +instance [semigroup β] [has_continuous_mul β] : semigroup C(α, β) := coe_injective.semigroup _ coe_mul @[to_additive] -instance {α : Type*} {β : Type*} [topological_space α] - [topological_space β] [comm_semigroup β] [has_continuous_mul β] : comm_semigroup C(α, β) := +instance [comm_semigroup β] [has_continuous_mul β] : comm_semigroup C(α, β) := coe_injective.comm_semigroup _ coe_mul @[to_additive] -instance {α : Type*} {β : Type*} [topological_space α] - [topological_space β] [mul_one_class β] [has_continuous_mul β] : mul_one_class C(α, β) := +instance [mul_one_class β] [has_continuous_mul β] : mul_one_class C(α, β) := coe_injective.mul_one_class _ coe_one coe_mul -instance {α : Type*} {β : Type*} [topological_space α] - [topological_space β] [mul_zero_class β] [has_continuous_mul β] : mul_zero_class C(α, β) := +instance [mul_zero_class β] [has_continuous_mul β] : mul_zero_class C(α, β) := coe_injective.mul_zero_class _ coe_zero coe_mul -instance {α : Type*} {β : Type*} [topological_space α] [topological_space β] - [semigroup_with_zero β] [has_continuous_mul β] : semigroup_with_zero C(α, β) := +instance [semigroup_with_zero β] [has_continuous_mul β] : semigroup_with_zero C(α, β) := coe_injective.semigroup_with_zero _ coe_zero coe_mul @[to_additive] -instance {α : Type*} {β : Type*} [topological_space α] [topological_space β] - [monoid β] [has_continuous_mul β] : monoid C(α, β) := +instance [monoid β] [has_continuous_mul β] : monoid C(α, β) := coe_injective.monoid _ coe_one coe_mul coe_pow -instance {α : Type*} {β : Type*} [topological_space α] [topological_space β] - [monoid_with_zero β] [has_continuous_mul β] : monoid_with_zero C(α, β) := +instance [monoid_with_zero β] [has_continuous_mul β] : monoid_with_zero C(α, β) := coe_injective.monoid_with_zero _ coe_zero coe_one coe_mul coe_pow @[to_additive] -instance {α : Type*} {β : Type*} [topological_space α] - [topological_space β] [comm_monoid β] [has_continuous_mul β] : comm_monoid C(α, β) := +instance [comm_monoid β] [has_continuous_mul β] : comm_monoid C(α, β) := coe_injective.comm_monoid _ coe_one coe_mul coe_pow -instance {α : Type*} {β : Type*} [topological_space α] [topological_space β] - [comm_monoid_with_zero β] [has_continuous_mul β] : comm_monoid_with_zero C(α, β) := +instance [comm_monoid_with_zero β] [has_continuous_mul β] : comm_monoid_with_zero C(α, β) := coe_injective.comm_monoid_with_zero _ coe_zero coe_one coe_mul coe_pow @[to_additive] -instance {α : Type*} {β : Type*} [topological_space α] - [locally_compact_space α] [topological_space β] - [has_mul β] [has_continuous_mul β] : has_continuous_mul C(α, β) := +instance [locally_compact_space α] [has_mul β] [has_continuous_mul β] : + has_continuous_mul C(α, β) := ⟨begin refine continuous_of_continuous_uncurry _ _, have h1 : continuous (λ x : (C(α, β) × C(α, β)) × α, x.fst.fst x.snd) := @@ -237,56 +230,53 @@ end⟩ /-- Coercion to a function as an `monoid_hom`. Similar to `monoid_hom.coe_fn`. -/ @[to_additive "Coercion to a function as an `add_monoid_hom`. Similar to `add_monoid_hom.coe_fn`.", simps] -def coe_fn_monoid_hom {α : Type*} {β : Type*} [topological_space α] [topological_space β] - [monoid β] [has_continuous_mul β] : C(α, β) →* (α → β) := +def coe_fn_monoid_hom [monoid β] [has_continuous_mul β] : C(α, β) →* (α → β) := { to_fun := coe_fn, map_one' := coe_one, map_mul' := coe_mul } +variables (α) + /-- Composition on the left by a (continuous) homomorphism of topological monoids, as a `monoid_hom`. Similar to `monoid_hom.comp_left`. -/ @[to_additive "Composition on the left by a (continuous) homomorphism of topological `add_monoid`s, as an `add_monoid_hom`. Similar to `add_monoid_hom.comp_left`.", simps] -protected def _root_.monoid_hom.comp_left_continuous (α : Type*) {β : Type*} {γ : Type*} - [topological_space α] [topological_space β] [monoid β] [has_continuous_mul β] +protected def _root_.monoid_hom.comp_left_continuous + {γ : Type*} [monoid β] [has_continuous_mul β] [topological_space γ] [monoid γ] [has_continuous_mul γ] (g : β →* γ) (hg : continuous g) : C(α, β) →* C(α, γ) := { to_fun := λ f, (⟨g, hg⟩ : C(β, γ)).comp f, map_one' := ext $ λ x, g.map_one, map_mul' := λ f₁ f₂, ext $ λ x, g.map_mul _ _ } +variables {α} + /-- Composition on the right as a `monoid_hom`. Similar to `monoid_hom.comp_hom'`. -/ @[to_additive "Composition on the right as an `add_monoid_hom`. Similar to `add_monoid_hom.comp_hom'`.", simps] -def comp_monoid_hom' {α : Type*} {β : Type*} {γ : Type*} - [topological_space α] [topological_space β] [topological_space γ] +def comp_monoid_hom' {γ : Type*} [topological_space γ] [mul_one_class γ] [has_continuous_mul γ] (g : C(α, β)) : C(β, γ) →* C(α, γ) := { to_fun := λ f, f.comp g, map_one' := one_comp g, map_mul' := λ f₁ f₂, mul_comp f₁ f₂ g } open_locale big_operators -@[simp, to_additive] lemma coe_prod {α : Type*} {β : Type*} [comm_monoid β] - [topological_space α] [topological_space β] [has_continuous_mul β] +@[simp, to_additive] lemma coe_prod [comm_monoid β] [has_continuous_mul β] {ι : Type*} (s : finset ι) (f : ι → C(α, β)) : ⇑(∏ i in s, f i) = (∏ i in s, (f i : α → β)) := (coe_fn_monoid_hom : C(α, β) →* _).map_prod f s @[to_additive] -lemma prod_apply {α : Type*} {β : Type*} [comm_monoid β] - [topological_space α] [topological_space β] [has_continuous_mul β] +lemma prod_apply [comm_monoid β] [has_continuous_mul β] {ι : Type*} (s : finset ι) (f : ι → C(α, β)) (a : α) : (∏ i in s, f i) a = (∏ i in s, f i a) := by simp @[to_additive] -instance {α : Type*} {β : Type*} [topological_space α] [topological_space β] - [group β] [topological_group β] : group C(α, β) := +instance [group β] [topological_group β] : group C(α, β) := coe_injective.group _ coe_one coe_mul coe_inv coe_div coe_pow coe_zpow @[to_additive] -instance {α : Type*} {β : Type*} [topological_space α] - [topological_space β] [comm_group β] [topological_group β] : comm_group C(α, β) := +instance [comm_group β] [topological_group β] : comm_group C(α, β) := coe_injective.comm_group _ coe_one coe_mul coe_inv coe_div coe_pow coe_zpow -@[to_additive] instance {α : Type*} {β : Type*} [topological_space α] - [topological_space β] [comm_group β] [topological_group β] : topological_group C(α, β) := +@[to_additive] instance [comm_group β] [topological_group β] : topological_group C(α, β) := { continuous_mul := by { letI : uniform_space β := topological_group.to_uniform_space β, have : uniform_group β := topological_comm_group_is_uniform, @@ -305,6 +295,30 @@ coe_injective.comm_group _ coe_one coe_mul coe_inv coe_div coe_pow coe_zpow exactI λ K hK, uniform_continuous_inv.comp_tendsto_uniformly_on (tendsto_iff_forall_compact_tendsto_uniformly_on.mp filter.tendsto_id K hK), } } +-- TODO: rewrite the next three lemmas for products and deduce sum case via `to_additive`, once +-- definition of `tprod` is in place + +/-- If `α` is locally compact, and an infinite sum of functions in `C(α, β)` +converges to `g` (for the compact-open topology), then the pointwise sum converges to `g x` for +all `x ∈ α`. -/ +lemma has_sum_apply {γ : Type*} [locally_compact_space α] [add_comm_monoid β] [has_continuous_add β] + {f : γ → C(α, β)} {g : C(α, β)} (hf : has_sum f g) (x : α) : + has_sum (λ i : γ, f i x) (g x) := +begin + let evₓ : add_monoid_hom C(α, β) β := (pi.eval_add_monoid_hom _ x).comp coe_fn_add_monoid_hom, + exact hf.map evₓ (continuous_map.continuous_eval_const' x), +end + +lemma summable_apply [locally_compact_space α] [add_comm_monoid β] [has_continuous_add β] + {γ : Type*} {f : γ → C(α, β)} (hf : summable f) (x : α) : + summable (λ i : γ, f i x) := +(has_sum_apply hf.has_sum x).summable + +lemma tsum_apply [locally_compact_space α] [t2_space β] [add_comm_monoid β] [has_continuous_add β] + {γ : Type*} {f : γ → C(α, β)} (hf : summable f) (x : α) : + (∑' (i:γ), f i x) = (∑' (i:γ), f i) x := +(has_sum_apply hf.has_sum x).tsum_eq + end continuous_map end group_structure From 71150516f28d9826c7341f8815b31f7d8770c212 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sat, 11 Feb 2023 10:50:05 +0000 Subject: [PATCH 0465/1291] chore(algebra/algebra/basic): remove a duplicate lemma (#18415) This combines `submodule.span_eq_restrict_scalars` (which was stated backwards and added in #6098) with `algebra.span_restrict_scalars_eq_span_of_surjective` (which was in an odd namespace, and added in #13042). The two lemmas were identical modulo argument order and explicitness. This also has the bonus of reducing the imports of `algebra.algebra.basic`. --- src/algebra/algebra/basic.lean | 22 ++------------ src/algebra/algebra/tower.lean | 32 ++++++++++----------- src/algebra/category/Module/tannaka.lean | 1 + src/number_theory/ramification_inertia.lean | 4 +-- src/ring_theory/derivation.lean | 4 +-- src/ring_theory/finiteness.lean | 2 +- 6 files changed, 23 insertions(+), 42 deletions(-) diff --git a/src/algebra/algebra/basic.lean b/src/algebra/algebra/basic.lean index b478de90292dc..9ce4bffda4eca 100644 --- a/src/algebra/algebra/basic.lean +++ b/src/algebra/algebra/basic.lean @@ -6,10 +6,11 @@ Authors: Kenny Lau, Yury Kudryashov import algebra.module.basic import algebra.module.ulift import algebra.ne_zero +import algebra.punit_instances import algebra.ring.aut import algebra.ring.ulift import algebra.char_zero.lemmas -import linear_algebra.span +import linear_algebra.basic import ring_theory.subring.basic import tactic.abel @@ -812,25 +813,6 @@ rfl end module -namespace submodule - -variables (R A M : Type*) -variables [comm_semiring R] [semiring A] [algebra R A] [add_comm_monoid M] -variables [module R M] [module A M] [is_scalar_tower R A M] - -/-- If `A` is an `R`-algebra such that the induced morhpsim `R →+* A` is surjective, then the -`R`-module generated by a set `X` equals the `A`-module generated by `X`. -/ -lemma span_eq_restrict_scalars (X : set M) (hsur : function.surjective (algebra_map R A)) : - span R X = restrict_scalars R (span A X) := -begin - apply (span_le_restrict_scalars R A X).antisymm (λ m hm, _), - refine span_induction hm subset_span (zero_mem _) (λ _ _, add_mem) (λ a m hm, _), - obtain ⟨r, rfl⟩ := hsur a, - simpa [algebra_map_smul] using smul_mem _ r hm -end - -end submodule - example {R A} [comm_semiring R] [semiring A] [module R A] [smul_comm_class R A A] [is_scalar_tower R A A] : algebra R A := algebra.of_module smul_mul_assoc mul_smul_comm diff --git a/src/algebra/algebra/tower.lean b/src/algebra/algebra/tower.lean index 20fd4850d1baa..5912d49fa0eb0 100644 --- a/src/algebra/algebra/tower.lean +++ b/src/algebra/algebra/tower.lean @@ -5,6 +5,7 @@ Authors: Kenny Lau, Anne Baanen -/ import algebra.algebra.equiv +import linear_algebra.span /-! # Towers of algebras @@ -190,32 +191,29 @@ end alg_equiv end homs -namespace algebra +namespace submodule -variables {R A} [comm_semiring R] [semiring A] [algebra R A] -variables {M} [add_comm_monoid M] [module A M] [module R M] [is_scalar_tower R A M] +variables (R A) {M} +variables [comm_semiring R] [semiring A] [algebra R A] [add_comm_monoid M] +variables [module R M] [module A M] [is_scalar_tower R A M] -lemma span_restrict_scalars_eq_span_of_surjective - (h : function.surjective (algebra_map R A)) (s : set M) : - (submodule.span A s).restrict_scalars R = submodule.span R s := +/-- If `A` is an `R`-algebra such that the induced morphism `R →+* A` is surjective, then the +`R`-module generated by a set `X` equals the `A`-module generated by `X`. -/ +lemma restrict_scalars_span (hsur : function.surjective (algebra_map R A)) (X : set M) : + restrict_scalars R (span A X) = span R X := begin - refine le_antisymm (λ x hx, _) (submodule.span_subset_span _ _ _), - refine submodule.span_induction hx _ _ _ _, - { exact λ x hx, submodule.subset_span hx }, - { exact submodule.zero_mem _ }, - { exact λ x y, submodule.add_mem _ }, - { intros c x hx, - obtain ⟨c', rfl⟩ := h c, - rw is_scalar_tower.algebra_map_smul, - exact submodule.smul_mem _ _ hx }, + refine ((span_le_restrict_scalars R A X).antisymm (λ m hm, _)).symm, + refine span_induction hm subset_span (zero_mem _) (λ _ _, add_mem) (λ a m hm, _), + obtain ⟨r, rfl⟩ := hsur a, + simpa [algebra_map_smul] using smul_mem _ r hm end lemma coe_span_eq_span_of_surjective (h : function.surjective (algebra_map R A)) (s : set M) : (submodule.span A s : set M) = submodule.span R s := -congr_arg coe (algebra.span_restrict_scalars_eq_span_of_surjective h s) +congr_arg coe (submodule.restrict_scalars_span R A h s) -end algebra +end submodule section semiring diff --git a/src/algebra/category/Module/tannaka.lean b/src/algebra/category/Module/tannaka.lean index fd59fa8f4bab3..4932673af4e7b 100644 --- a/src/algebra/category/Module/tannaka.lean +++ b/src/algebra/category/Module/tannaka.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ import algebra.category.Module.basic +import linear_algebra.span /-! # Tannaka duality for rings diff --git a/src/number_theory/ramification_inertia.lean b/src/number_theory/ramification_inertia.lean index 4c77e129f0147..38a4402e6e9bb 100644 --- a/src/number_theory/ramification_inertia.lean +++ b/src/number_theory/ramification_inertia.lean @@ -417,8 +417,8 @@ begin have mem_span_b : ((submodule.mkq (map (algebra_map R S) p)) x : S ⧸ map (algebra_map R S) p) ∈ submodule.span (R ⧸ p) (set.range b) := b.mem_span _, - rw [← @submodule.restrict_scalars_mem R, algebra.span_restrict_scalars_eq_span_of_surjective - (show function.surjective (algebra_map R (R ⧸ p)), from ideal.quotient.mk_surjective) _, + rw [← @submodule.restrict_scalars_mem R, + submodule.restrict_scalars_span R (R ⧸ p) ideal.quotient.mk_surjective, b_eq_b', set.range_comp, ← submodule.map_span] at mem_span_b, obtain ⟨y, y_mem, y_eq⟩ := submodule.mem_map.mp mem_span_b, diff --git a/src/ring_theory/derivation.lean b/src/ring_theory/derivation.lean index 8f050414aca68..940fe5076a30f 100644 --- a/src/ring_theory/derivation.lean +++ b/src/ring_theory/derivation.lean @@ -1004,7 +1004,7 @@ lemma kaehler_differential.ker_total_map (h : function.surjective (algebra_map A (kaehler_differential.ker_total S B).restrict_scalars _ := begin rw [kaehler_differential.ker_total, submodule.map_span, kaehler_differential.ker_total, - ← submodule.span_eq_restrict_scalars _ _ _ _ h], + submodule.restrict_scalars_span _ _ h], simp_rw [set.image_union, submodule.span_union, ← set.image_univ, set.image_image, set.image_univ, map_sub, map_add], simp only [linear_map.comp_apply, finsupp.map_range.linear_map_apply, finsupp.map_range_single, @@ -1070,7 +1070,7 @@ lemma kaehler_differential.map_surjective_of_surjective function.surjective (kaehler_differential.map R S A B) := begin rw [← linear_map.range_eq_top, _root_.eq_top_iff, ← @submodule.restrict_scalars_top B A, - ← kaehler_differential.span_range_derivation, ← submodule.span_eq_restrict_scalars _ _ _ _ h, + ← kaehler_differential.span_range_derivation, submodule.restrict_scalars_span _ _ h, submodule.span_le], rintros _ ⟨x, rfl⟩, obtain ⟨y, rfl⟩ := h x, diff --git a/src/ring_theory/finiteness.lean b/src/ring_theory/finiteness.lean index 5ab0cc4e0c2ef..620942701459e 100644 --- a/src/ring_theory/finiteness.lean +++ b/src/ring_theory/finiteness.lean @@ -284,7 +284,7 @@ lemma fg_restrict_scalars {R S M : Type*} [comm_semiring R] [semiring S] [algebr begin obtain ⟨X, rfl⟩ := hfin, use X, - exact submodule.span_eq_restrict_scalars R S M X h + exact (submodule.restrict_scalars_span R S h ↑X).symm end lemma fg.stablizes_of_supr_eq {M' : submodule R M} (hM' : M'.fg) From 99624c1a1c3d353ed60e0a13e03246ce4807b56d Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Sat, 11 Feb 2023 16:35:52 +0000 Subject: [PATCH 0466/1291] chore(data/finset/sym): fix a typo (#18426) During the port of this file to mathlib4 (see https://github.com/leanprover-community/mathlib4/pull/2168#), it was noticed that the same alias was declared twice. Looking at the code, it is clear that the second declaration is a typo that it is fixed in this PR. --- src/data/finset/sym.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/data/finset/sym.lean b/src/data/finset/sym.lean index 408abf97802cf..bed5ded31f54b 100644 --- a/src/data/finset/sym.lean +++ b/src/data/finset/sym.lean @@ -141,9 +141,9 @@ end @[simp] lemma sym_nonempty : (s.sym n).nonempty ↔ n = 0 ∨ s.nonempty := by simp_rw [nonempty_iff_ne_empty, ne.def, sym_eq_empty, not_and_distrib, not_ne_iff] -alias sym2_nonempty ↔ _ nonempty.sym2 +alias sym_nonempty ↔ _ nonempty.sym -attribute [protected] nonempty.sym2 +attribute [protected] nonempty.sym @[simp] lemma sym_univ [fintype α] (n : ℕ) : (univ : finset α).sym n = univ := eq_univ_iff_forall.2 $ λ s, mem_sym_iff.2 $ λ a _, mem_univ _ From 70a4f2197832bceab57d7f41379b2592d1110570 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sun, 12 Feb 2023 08:45:16 +0000 Subject: [PATCH 0467/1291] =?UTF-8?q?refactor(measure=5Ftheory/measure/mut?= =?UTF-8?q?ually=5Fsingular):=20use=20=E2=9F=82=20PERPENDICULAR=20instead?= =?UTF-8?q?=20of=20=E2=8A=A5=20UP=20TACK=20(#18423)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The previous notation for `measure_theory.measure.mutually_singular` was `⊥ₘ`. This changes it to `⟂ₘ`, which is semantically a better character. The same change is made for `measure_theory.vector_measure.mutually_singular`, from `⊥ᵥ` to `⟂ᵥ`. [Zulip](https://leanprover.zulipchat.com/#narrow/stream/113488-general/topic/Perpendicular.20notation/near/327158463) --- .../covering/differentiation.lean | 2 +- src/measure_theory/decomposition/jordan.lean | 10 ++-- .../decomposition/lebesgue.lean | 48 +++++++++---------- src/measure_theory/integral/lebesgue.lean | 2 +- .../measure/mutually_singular.lean | 34 ++++++------- .../measure/vector_measure.lean | 26 +++++----- 6 files changed, 61 insertions(+), 61 deletions(-) diff --git a/src/measure_theory/covering/differentiation.lean b/src/measure_theory/covering/differentiation.lean index ba4d4db70ca10..7aa3813ee22a3 100644 --- a/src/measure_theory/covering/differentiation.lean +++ b/src/measure_theory/covering/differentiation.lean @@ -158,7 +158,7 @@ variables [second_countable_topology α] [borel_space α] [is_locally_finite_mea /-- If a measure `ρ` is singular with respect to `μ`, then for `μ` almost every `x`, the ratio `ρ a / μ a` tends to zero when `a` shrinks to `x` along the Vitali family. This makes sense as `μ a` is eventually positive by `ae_eventually_measure_pos`. -/ -lemma ae_eventually_measure_zero_of_singular (hρ : ρ ⊥ₘ μ) : +lemma ae_eventually_measure_zero_of_singular (hρ : ρ ⟂ₘ μ) : ∀ᵐ x ∂μ, tendsto (λ a, ρ a / μ a) (v.filter_at x) (𝓝 0) := begin have A : ∀ ε > (0 : ℝ≥0), ∀ᵐ x ∂μ, ∀ᶠ a in v.filter_at x, ρ a < ε * μ a, diff --git a/src/measure_theory/decomposition/jordan.lean b/src/measure_theory/decomposition/jordan.lean index a332a7ca18a72..de697ffb6d442 100644 --- a/src/measure_theory/decomposition/jordan.lean +++ b/src/measure_theory/decomposition/jordan.lean @@ -52,7 +52,7 @@ finite measures. -/ (pos_part neg_part : measure α) [pos_part_finite : is_finite_measure pos_part] [neg_part_finite : is_finite_measure neg_part] -(mutually_singular : pos_part ⊥ₘ neg_part) +(mutually_singular : pos_part ⟂ₘ neg_part) attribute [instance] jordan_decomposition.pos_part_finite attribute [instance] jordan_decomposition.neg_part_finite @@ -510,7 +510,7 @@ end -- TODO: Generalize to vector measures once total variation on vector measures is defined lemma mutually_singular_iff (s t : signed_measure α) : - s ⊥ᵥ t ↔ s.total_variation ⊥ₘ t.total_variation := + s ⟂ᵥ t ↔ s.total_variation ⟂ₘ t.total_variation := begin split, { rintro ⟨u, hmeas, hu₁, hu₂⟩, @@ -531,7 +531,7 @@ begin end lemma mutually_singular_ennreal_iff (s : signed_measure α) (μ : vector_measure α ℝ≥0∞) : - s ⊥ᵥ μ ↔ s.total_variation ⊥ₘ μ.ennreal_to_measure := + s ⟂ᵥ μ ↔ s.total_variation ⟂ₘ μ.ennreal_to_measure := begin split, { rintro ⟨u, hmeas, hu₁, hu₂⟩, @@ -550,8 +550,8 @@ begin end lemma total_variation_mutually_singular_iff (s : signed_measure α) (μ : measure α) : - s.total_variation ⊥ₘ μ ↔ - s.to_jordan_decomposition.pos_part ⊥ₘ μ ∧ s.to_jordan_decomposition.neg_part ⊥ₘ μ := + s.total_variation ⟂ₘ μ ↔ + s.to_jordan_decomposition.pos_part ⟂ₘ μ ∧ s.to_jordan_decomposition.neg_part ⟂ₘ μ := measure.mutually_singular.add_left_iff end signed_measure diff --git a/src/measure_theory/decomposition/lebesgue.lean b/src/measure_theory/decomposition/lebesgue.lean index f380e3101d167..85b42905182f4 100644 --- a/src/measure_theory/decomposition/lebesgue.lean +++ b/src/measure_theory/decomposition/lebesgue.lean @@ -74,7 +74,7 @@ measure `ξ` and a measurable function `f`, such that `ξ` is mutually singular `ν` and `μ = ξ + ν.with_density f`. -/ class have_lebesgue_decomposition (μ ν : measure α) : Prop := (lebesgue_decomposition : - ∃ (p : measure α × (α → ℝ≥0∞)), measurable p.2 ∧ p.1 ⊥ₘ ν ∧ μ = p.1 + ν.with_density p.2) + ∃ (p : measure α × (α → ℝ≥0∞)), measurable p.2 ∧ p.1 ⟂ₘ ν ∧ μ = p.1 + ν.with_density p.2) /-- If a pair of measures `have_lebesgue_decomposition`, then `singular_part` chooses the measure from `have_lebesgue_decomposition`, otherwise it returns the zero measure. For sigma-finite @@ -92,7 +92,7 @@ if h : have_lebesgue_decomposition μ ν then (classical.some h.lebesgue_decompo lemma have_lebesgue_decomposition_spec (μ ν : measure α) [h : have_lebesgue_decomposition μ ν] : - measurable (μ.rn_deriv ν) ∧ (μ.singular_part ν) ⊥ₘ ν ∧ + measurable (μ.rn_deriv ν) ∧ (μ.singular_part ν) ⟂ₘ ν ∧ μ = (μ.singular_part ν) + ν.with_density (μ.rn_deriv ν) := begin rw [singular_part, rn_deriv, dif_pos h, dif_pos h], @@ -129,7 +129,7 @@ begin end lemma mutually_singular_singular_part (μ ν : measure α) : - μ.singular_part ν ⊥ₘ ν := + μ.singular_part ν ⟂ₘ ν := begin by_cases h : have_lebesgue_decomposition μ ν, { exactI (have_lebesgue_decomposition_spec μ ν).2.1 }, @@ -226,7 +226,7 @@ This theorem provides the uniqueness of the `singular_part` in the Lebesgue deco while `measure_theory.measure.eq_rn_deriv` provides the uniqueness of the `rn_deriv`. -/ theorem eq_singular_part {s : measure α} {f : α → ℝ≥0∞} (hf : measurable f) - (hs : s ⊥ₘ ν) (hadd : μ = s + ν.with_density f) : + (hs : s ⟂ₘ ν) (hadd : μ = s + ν.with_density f) : s = μ.singular_part ν := begin haveI : have_lebesgue_decomposition μ ν := ⟨⟨⟨s, f⟩, hf, hs, hadd⟩⟩, @@ -316,7 +316,7 @@ theorem, while `measure_theory.measure.eq_singular_part` provides the uniqueness `singular_part`. Here, the uniqueness is given in terms of the measures, while the uniqueness in terms of the functions is given in `eq_rn_deriv`. -/ theorem eq_with_density_rn_deriv {s : measure α} {f : α → ℝ≥0∞} (hf : measurable f) - (hs : s ⊥ₘ ν) (hadd : μ = s + ν.with_density f) : + (hs : s ⟂ₘ ν) (hadd : μ = s + ν.with_density f) : ν.with_density f = ν.with_density (μ.rn_deriv ν) := begin haveI : have_lebesgue_decomposition μ ν := ⟨⟨⟨s, f⟩, hf, hs, hadd⟩⟩, @@ -366,7 +366,7 @@ theorem, while `measure_theory.measure.eq_singular_part` provides the uniqueness `singular_part`. Here, the uniqueness is given in terms of the functions, while the uniqueness in terms of the functions is given in `eq_with_density_rn_deriv`. -/ theorem eq_rn_deriv [sigma_finite ν] {s : measure α} {f : α → ℝ≥0∞} (hf : measurable f) - (hs : s ⊥ₘ ν) (hadd : μ = s + ν.with_density f) : + (hs : s ⟂ₘ ν) (hadd : μ = s + ν.with_density f) : f =ᵐ[ν] μ.rn_deriv ν := begin refine ae_eq_of_forall_set_lintegral_eq_of_sigma_finite hf (measurable_rn_deriv μ ν) _, @@ -400,7 +400,7 @@ a measurable set `E`, such that `ν(E) > 0` and `E` is positive with respect to This lemma is useful for the Lebesgue decomposition theorem. -/ lemma exists_positive_of_not_mutually_singular - (μ ν : measure α) [is_finite_measure μ] [is_finite_measure ν] (h : ¬ μ ⊥ₘ ν) : + (μ ν : measure α) [is_finite_measure μ] [is_finite_measure ν] (h : ¬ μ ⟂ₘ ν) : ∃ ε : ℝ≥0, 0 < ε ∧ ∃ E : set α, measurable_set E ∧ 0 < ν E ∧ 0 ≤[E] μ.to_signed_measure - (ε • ν).to_signed_measure := begin @@ -704,7 +704,7 @@ instance have_lebesgue_decomposition_of_sigma_finite { choose A hA₁ hA₂ hA₃ using λ n, mutually_singular_singular_part (μn n) (νn n), simp only [hξ], -- We use the set `B := ⋃ j, (S.set j) ∩ A j` where `A n` is the set provided as - -- `singular_part (μn n) (νn n) ⊥ₘ νn n` + -- `singular_part (μn n) (νn n) ⟂ₘ νn n` refine ⟨⋃ j, (S.set j) ∩ A j, measurable_set.Union (λ n, (S.set_mem n).inter (hA₁ n)), _, _⟩, -- `ξ B = 0` since `ξ B = ∑ i j, singular_part (μn j) (νn j) (S.set i ∩ A i)` @@ -855,7 +855,7 @@ def singular_part (s : signed_measure α) (μ : measure α) : signed_measure α section lemma singular_part_mutually_singular (s : signed_measure α) (μ : measure α) : - s.to_jordan_decomposition.pos_part.singular_part μ ⊥ₘ + s.to_jordan_decomposition.pos_part.singular_part μ ⟂ₘ s.to_jordan_decomposition.neg_part.singular_part μ := begin by_cases hl : s.have_lebesgue_decomposition μ, @@ -888,10 +888,10 @@ begin end lemma mutually_singular_singular_part (s : signed_measure α) (μ : measure α) : - singular_part s μ ⊥ᵥ μ.to_ennreal_vector_measure := + singular_part s μ ⟂ᵥ μ.to_ennreal_vector_measure := begin rw [mutually_singular_ennreal_iff, singular_part_total_variation], - change _ ⊥ₘ vector_measure.equiv_measure.to_fun (vector_measure.equiv_measure.inv_fun μ), + change _ ⟂ₘ vector_measure.equiv_measure.to_fun (vector_measure.equiv_measure.inv_fun μ), rw vector_measure.equiv_measure.right_inv μ, exact (mutually_singular_singular_part _ _).add_left (mutually_singular_singular_part _ _) end @@ -958,13 +958,13 @@ end variables {s μ} lemma jordan_decomposition_add_with_density_mutually_singular - {f : α → ℝ} (hf : measurable f) (htμ : t ⊥ᵥ μ.to_ennreal_vector_measure) : - t.to_jordan_decomposition.pos_part + μ.with_density (λ (x : α), ennreal.of_real (f x)) ⊥ₘ + {f : α → ℝ} (hf : measurable f) (htμ : t ⟂ᵥ μ.to_ennreal_vector_measure) : + t.to_jordan_decomposition.pos_part + μ.with_density (λ (x : α), ennreal.of_real (f x)) ⟂ₘ t.to_jordan_decomposition.neg_part + μ.with_density (λ (x : α), ennreal.of_real (-f x)) := begin rw [mutually_singular_ennreal_iff, total_variation_mutually_singular_iff] at htμ, - change _ ⊥ₘ vector_measure.equiv_measure.to_fun (vector_measure.equiv_measure.inv_fun μ) ∧ - _ ⊥ₘ vector_measure.equiv_measure.to_fun (vector_measure.equiv_measure.inv_fun μ) at htμ, + change _ ⟂ₘ vector_measure.equiv_measure.to_fun (vector_measure.equiv_measure.inv_fun μ) ∧ + _ ⟂ₘ vector_measure.equiv_measure.to_fun (vector_measure.equiv_measure.inv_fun μ) at htμ, rw [vector_measure.equiv_measure.right_inv] at htμ, exact ((jordan_decomposition.mutually_singular _).add_right (htμ.1.mono_ac (refl _) (with_density_absolutely_continuous _ _))).add_left @@ -974,7 +974,7 @@ end lemma to_jordan_decomposition_eq_of_eq_add_with_density {f : α → ℝ} (hf : measurable f) (hfi : integrable f μ) - (htμ : t ⊥ᵥ μ.to_ennreal_vector_measure) (hadd : s = t + μ.with_densityᵥ f) : + (htμ : t ⟂ᵥ μ.to_ennreal_vector_measure) (hadd : s = t + μ.with_densityᵥ f) : s.to_jordan_decomposition = @jordan_decomposition.mk α _ (t.to_jordan_decomposition.pos_part + μ.with_density (λ x, ennreal.of_real (f x))) (t.to_jordan_decomposition.neg_part + μ.with_density (λ x, ennreal.of_real (- f x))) @@ -1001,12 +1001,12 @@ end private lemma have_lebesgue_decomposition_mk' (μ : measure α) {f : α → ℝ} (hf : measurable f) (hfi : integrable f μ) - (htμ : t ⊥ᵥ μ.to_ennreal_vector_measure) (hadd : s = t + μ.with_densityᵥ f) : + (htμ : t ⟂ᵥ μ.to_ennreal_vector_measure) (hadd : s = t + μ.with_densityᵥ f) : s.have_lebesgue_decomposition μ := begin have htμ' := htμ, rw mutually_singular_ennreal_iff at htμ, - change _ ⊥ₘ vector_measure.equiv_measure.to_fun (vector_measure.equiv_measure.inv_fun μ) at htμ, + change _ ⟂ₘ vector_measure.equiv_measure.to_fun (vector_measure.equiv_measure.inv_fun μ) at htμ, rw [vector_measure.equiv_measure.right_inv, total_variation_mutually_singular_iff] at htμ, refine { pos_part := @@ -1020,7 +1020,7 @@ begin end lemma have_lebesgue_decomposition_mk (μ : measure α) {f : α → ℝ} (hf : measurable f) - (htμ : t ⊥ᵥ μ.to_ennreal_vector_measure) (hadd : s = t + μ.with_densityᵥ f) : + (htμ : t ⟂ᵥ μ.to_ennreal_vector_measure) (hadd : s = t + μ.with_densityᵥ f) : s.have_lebesgue_decomposition μ := begin by_cases hfi : integrable f μ, @@ -1032,13 +1032,13 @@ end private theorem eq_singular_part' (t : signed_measure α) {f : α → ℝ} (hf : measurable f) (hfi : integrable f μ) - (htμ : t ⊥ᵥ μ.to_ennreal_vector_measure) (hadd : s = t + μ.with_densityᵥ f) : + (htμ : t ⟂ᵥ μ.to_ennreal_vector_measure) (hadd : s = t + μ.with_densityᵥ f) : t = s.singular_part μ := begin have htμ' := htμ, rw [mutually_singular_ennreal_iff, total_variation_mutually_singular_iff] at htμ, - change _ ⊥ₘ vector_measure.equiv_measure.to_fun (vector_measure.equiv_measure.inv_fun μ) ∧ - _ ⊥ₘ vector_measure.equiv_measure.to_fun (vector_measure.equiv_measure.inv_fun μ) at htμ, + change _ ⟂ₘ vector_measure.equiv_measure.to_fun (vector_measure.equiv_measure.inv_fun μ) ∧ + _ ⟂ₘ vector_measure.equiv_measure.to_fun (vector_measure.equiv_measure.inv_fun μ) at htμ, rw [vector_measure.equiv_measure.right_inv] at htμ, { rw [singular_part, ← t.to_signed_measure_to_jordan_decomposition, jordan_decomposition.to_signed_measure], @@ -1056,7 +1056,7 @@ mutually singular with respect to `μ` and `s = t + μ.with_densityᵥ f`, we ha `t = singular_part s μ`, i.e. `t` is the singular part of the Lebesgue decomposition between `s` and `μ`. -/ theorem eq_singular_part (t : signed_measure α) (f : α → ℝ) - (htμ : t ⊥ᵥ μ.to_ennreal_vector_measure) (hadd : s = t + μ.with_densityᵥ f) : + (htμ : t ⟂ᵥ μ.to_ennreal_vector_measure) (hadd : s = t + μ.with_densityᵥ f) : t = s.singular_part μ := begin by_cases hfi : integrable f μ, @@ -1139,7 +1139,7 @@ by { rw [sub_eq_add_neg, sub_eq_add_neg, singular_part_add, singular_part_neg] } mutually singular with respect to `μ` and `s = t + μ.with_densityᵥ f`, we have `f = rn_deriv s μ`, i.e. `f` is the Radon-Nikodym derivative of `s` and `μ`. -/ theorem eq_rn_deriv (t : signed_measure α) (f : α → ℝ) (hfi : integrable f μ) - (htμ : t ⊥ᵥ μ.to_ennreal_vector_measure) (hadd : s = t + μ.with_densityᵥ f) : + (htμ : t ⟂ᵥ μ.to_ennreal_vector_measure) (hadd : s = t + μ.with_densityᵥ f) : f =ᵐ[μ] s.rn_deriv μ := begin set f' := hfi.1.mk f, diff --git a/src/measure_theory/integral/lebesgue.lean b/src/measure_theory/integral/lebesgue.lean index efb3bed88cd4f..f90e0922cd070 100644 --- a/src/measure_theory/integral/lebesgue.lean +++ b/src/measure_theory/integral/lebesgue.lean @@ -2483,7 +2483,7 @@ lemma with_density_indicator_one {s : set α} (hs : measurable_set s) : by rw [with_density_indicator hs, with_density_one] lemma with_density_of_real_mutually_singular {f : α → ℝ} (hf : measurable f) : - μ.with_density (λ x, ennreal.of_real $ f x) ⊥ₘ μ.with_density (λ x, ennreal.of_real $ -f x) := + μ.with_density (λ x, ennreal.of_real $ f x) ⟂ₘ μ.with_density (λ x, ennreal.of_real $ -f x) := begin set S : set α := { x | f x < 0 } with hSdef, have hS : measurable_set S := measurable_set_lt hf measurable_const, diff --git a/src/measure_theory/measure/mutually_singular.lean b/src/measure_theory/measure/mutually_singular.lean index 33b5e4ce06461..0146f97cf9d92 100644 --- a/src/measure_theory/measure/mutually_singular.lean +++ b/src/measure_theory/measure/mutually_singular.lean @@ -8,9 +8,9 @@ import measure_theory.measure.measure_space /-! # Mutually singular measures Two measures `μ`, `ν` are said to be mutually singular (`measure_theory.measure.mutually_singular`, -localized notation `μ ⊥ₘ ν`) if there exists a measurable set `s` such that `μ s = 0` and +localized notation `μ ⟂ₘ ν`) if there exists a measurable set `s` such that `μ s = 0` and `ν sᶜ = 0`. The measurability of `s` is an unnecessary assumption (see -`measure_theory.measure.mutually_singular.mk`) but we keep it because this way `rcases (h : μ ⊥ₘ ν)` +`measure_theory.measure.mutually_singular.mk`) but we keep it because this way `rcases (h : μ ⟂ₘ ν)` gives us a measurable set and usually it is easy to prove measurability. In this file we define the predicate `measure_theory.measure.mutually_singular` and prove basic @@ -36,7 +36,7 @@ def mutually_singular {m0 : measurable_space α} (μ ν : measure α) : Prop := ∃ (s : set α), measurable_set s ∧ μ s = 0 ∧ ν sᶜ = 0 localized "infix (name := measure.mutually_singular) - ` ⊥ₘ `:60 := measure_theory.measure.mutually_singular" in measure_theory + ` ⟂ₘ `:60 := measure_theory.measure.mutually_singular" in measure_theory namespace mutually_singular @@ -48,23 +48,23 @@ begin exact subset_to_measurable _ _ hxs end -@[simp] lemma zero_right : μ ⊥ₘ 0 := ⟨∅, measurable_set.empty, measure_empty, rfl⟩ +@[simp] lemma zero_right : μ ⟂ₘ 0 := ⟨∅, measurable_set.empty, measure_empty, rfl⟩ -@[symm] lemma symm (h : ν ⊥ₘ μ) : μ ⊥ₘ ν := +@[symm] lemma symm (h : ν ⟂ₘ μ) : μ ⟂ₘ ν := let ⟨i, hi, his, hit⟩ := h in ⟨iᶜ, hi.compl, hit, (compl_compl i).symm ▸ his⟩ -lemma comm : μ ⊥ₘ ν ↔ ν ⊥ₘ μ := ⟨λ h, h.symm, λ h, h.symm⟩ +lemma comm : μ ⟂ₘ ν ↔ ν ⟂ₘ μ := ⟨λ h, h.symm, λ h, h.symm⟩ -@[simp] lemma zero_left : 0 ⊥ₘ μ := zero_right.symm +@[simp] lemma zero_left : 0 ⟂ₘ μ := zero_right.symm -lemma mono_ac (h : μ₁ ⊥ₘ ν₁) (hμ : μ₂ ≪ μ₁) (hν : ν₂ ≪ ν₁) : μ₂ ⊥ₘ ν₂ := +lemma mono_ac (h : μ₁ ⟂ₘ ν₁) (hμ : μ₂ ≪ μ₁) (hν : ν₂ ≪ ν₁) : μ₂ ⟂ₘ ν₂ := let ⟨s, hs, h₁, h₂⟩ := h in ⟨s, hs, hμ h₁, hν h₂⟩ -lemma mono (h : μ₁ ⊥ₘ ν₁) (hμ : μ₂ ≤ μ₁) (hν : ν₂ ≤ ν₁) : μ₂ ⊥ₘ ν₂ := +lemma mono (h : μ₁ ⟂ₘ ν₁) (hμ : μ₂ ≤ μ₁) (hν : ν₂ ≤ ν₁) : μ₂ ⟂ₘ ν₂ := h.mono_ac hμ.absolutely_continuous hν.absolutely_continuous @[simp] lemma sum_left {ι : Type*} [countable ι] {μ : ι → measure α} : - (sum μ) ⊥ₘ ν ↔ ∀ i, μ i ⊥ₘ ν := + (sum μ) ⟂ₘ ν ↔ ∀ i, μ i ⟂ₘ ν := begin refine ⟨λ h i, h.mono (le_sum _ _) le_rfl, λ H, _⟩, choose s hsm hsμ hsν using H, @@ -75,25 +75,25 @@ begin end @[simp] lemma sum_right {ι : Type*} [countable ι] {ν : ι → measure α} : - μ ⊥ₘ sum ν ↔ ∀ i, μ ⊥ₘ ν i := + μ ⟂ₘ sum ν ↔ ∀ i, μ ⟂ₘ ν i := comm.trans $ sum_left.trans $ forall_congr $ λ i, comm -@[simp] lemma add_left_iff : μ₁ + μ₂ ⊥ₘ ν ↔ μ₁ ⊥ₘ ν ∧ μ₂ ⊥ₘ ν := +@[simp] lemma add_left_iff : μ₁ + μ₂ ⟂ₘ ν ↔ μ₁ ⟂ₘ ν ∧ μ₂ ⟂ₘ ν := by rw [← sum_cond, sum_left, bool.forall_bool, cond, cond, and.comm] -@[simp] lemma add_right_iff : μ ⊥ₘ ν₁ + ν₂ ↔ μ ⊥ₘ ν₁ ∧ μ ⊥ₘ ν₂ := +@[simp] lemma add_right_iff : μ ⟂ₘ ν₁ + ν₂ ↔ μ ⟂ₘ ν₁ ∧ μ ⟂ₘ ν₂ := comm.trans $ add_left_iff.trans $ and_congr comm comm -lemma add_left (h₁ : ν₁ ⊥ₘ μ) (h₂ : ν₂ ⊥ₘ μ) : ν₁ + ν₂ ⊥ₘ μ := +lemma add_left (h₁ : ν₁ ⟂ₘ μ) (h₂ : ν₂ ⟂ₘ μ) : ν₁ + ν₂ ⟂ₘ μ := add_left_iff.2 ⟨h₁, h₂⟩ -lemma add_right (h₁ : μ ⊥ₘ ν₁) (h₂ : μ ⊥ₘ ν₂) : μ ⊥ₘ ν₁ + ν₂ := +lemma add_right (h₁ : μ ⟂ₘ ν₁) (h₂ : μ ⟂ₘ ν₂) : μ ⟂ₘ ν₁ + ν₂ := add_right_iff.2 ⟨h₁, h₂⟩ -lemma smul (r : ℝ≥0∞) (h : ν ⊥ₘ μ) : r • ν ⊥ₘ μ := +lemma smul (r : ℝ≥0∞) (h : ν ⟂ₘ μ) : r • ν ⟂ₘ μ := h.mono_ac (absolutely_continuous.rfl.smul r) absolutely_continuous.rfl -lemma smul_nnreal (r : ℝ≥0) (h : ν ⊥ₘ μ) : r • ν ⊥ₘ μ := h.smul r +lemma smul_nnreal (r : ℝ≥0) (h : ν ⟂ₘ μ) : r • ν ⟂ₘ μ := h.smul r end mutually_singular diff --git a/src/measure_theory/measure/vector_measure.lean b/src/measure_theory/measure/vector_measure.lean index 3bce12047f068..f29742363b51f 100644 --- a/src/measure_theory/measure/vector_measure.lean +++ b/src/measure_theory/measure/vector_measure.lean @@ -1103,7 +1103,7 @@ def mutually_singular (v : vector_measure α M) (w : vector_measure α N) : Prop ∃ (s : set α), measurable_set s ∧ (∀ t ⊆ s, v t = 0) ∧ (∀ t ⊆ sᶜ, w t = 0) localized "infix (name := vector_measure.mutually_singular) - ` ⊥ᵥ `:60 := measure_theory.vector_measure.mutually_singular" in measure_theory + ` ⟂ᵥ `:60 := measure_theory.vector_measure.mutually_singular" in measure_theory namespace mutually_singular @@ -1111,7 +1111,7 @@ variables {v v₁ v₂ : vector_measure α M} {w w₁ w₂ : vector_measure α N lemma mk (s : set α) (hs : measurable_set s) (h₁ : ∀ t ⊆ s, measurable_set t → v t = 0) - (h₂ : ∀ t ⊆ sᶜ, measurable_set t → w t = 0) : v ⊥ᵥ w := + (h₂ : ∀ t ⊆ sᶜ, measurable_set t → w t = 0) : v ⟂ᵥ w := begin refine ⟨s, hs, λ t hst, _, λ t hst, _⟩; by_cases ht : measurable_set t, @@ -1121,17 +1121,17 @@ begin { exact not_measurable w ht } end -lemma symm (h : v ⊥ᵥ w) : w ⊥ᵥ v := +lemma symm (h : v ⟂ᵥ w) : w ⟂ᵥ v := let ⟨s, hmeas, hs₁, hs₂⟩ := h in ⟨sᶜ, hmeas.compl, hs₂, λ t ht, hs₁ _ (compl_compl s ▸ ht : t ⊆ s)⟩ -lemma zero_right : v ⊥ᵥ (0 : vector_measure α N) := +lemma zero_right : v ⟂ᵥ (0 : vector_measure α N) := ⟨∅, measurable_set.empty, λ t ht, (subset_empty_iff.1 ht).symm ▸ v.empty, λ _ _, zero_apply _⟩ -lemma zero_left : (0 : vector_measure α M) ⊥ᵥ w := +lemma zero_left : (0 : vector_measure α M) ⟂ᵥ w := zero_right.symm -lemma add_left [t2_space N] [has_continuous_add M] (h₁ : v₁ ⊥ᵥ w) (h₂ : v₂ ⊥ᵥ w) : v₁ + v₂ ⊥ᵥ w := +lemma add_left [t2_space N] [has_continuous_add M] (h₁ : v₁ ⟂ᵥ w) (h₂ : v₂ ⟂ᵥ w) : v₁ + v₂ ⟂ᵥ w := begin obtain ⟨u, hmu, hu₁, hu₂⟩ := h₁, obtain ⟨v, hmv, hv₁, hv₂⟩ := h₂, @@ -1154,20 +1154,20 @@ begin { rcases hx; exact hx.2 } } }, end -lemma add_right [t2_space M] [has_continuous_add N] (h₁ : v ⊥ᵥ w₁) (h₂ : v ⊥ᵥ w₂) : v ⊥ᵥ w₁ + w₂ := +lemma add_right [t2_space M] [has_continuous_add N] (h₁ : v ⟂ᵥ w₁) (h₂ : v ⟂ᵥ w₂) : v ⟂ᵥ w₁ + w₂ := (add_left h₁.symm h₂.symm).symm lemma smul_right {R : Type*} [semiring R] [distrib_mul_action R N] [has_continuous_const_smul R N] - (r : R) (h : v ⊥ᵥ w) : v ⊥ᵥ r • w := + (r : R) (h : v ⟂ᵥ w) : v ⟂ᵥ r • w := let ⟨s, hmeas, hs₁, hs₂⟩ := h in ⟨s, hmeas, hs₁, λ t ht, by simp only [coe_smul, pi.smul_apply, hs₂ t ht, smul_zero]⟩ lemma smul_left {R : Type*} [semiring R] [distrib_mul_action R M] [has_continuous_const_smul R M] - (r : R) (h : v ⊥ᵥ w) : r • v ⊥ᵥ w := + (r : R) (h : v ⟂ᵥ w) : r • v ⟂ᵥ w := (smul_right r h.symm).symm lemma neg_left {M : Type*} [add_comm_group M] [topological_space M] [topological_add_group M] - {v : vector_measure α M} {w : vector_measure α N} (h : v ⊥ᵥ w) : -v ⊥ᵥ w := + {v : vector_measure α M} {w : vector_measure α N} (h : v ⟂ᵥ w) : -v ⟂ᵥ w := begin obtain ⟨u, hmu, hu₁, hu₂⟩ := h, refine ⟨u, hmu, λ s hs, _, hu₂⟩, @@ -1176,19 +1176,19 @@ begin end lemma neg_right {N : Type*} [add_comm_group N] [topological_space N] [topological_add_group N] - {v : vector_measure α M} {w : vector_measure α N} (h : v ⊥ᵥ w) : v ⊥ᵥ -w := + {v : vector_measure α M} {w : vector_measure α N} (h : v ⟂ᵥ w) : v ⟂ᵥ -w := h.symm.neg_left.symm @[simp] lemma neg_left_iff {M : Type*} [add_comm_group M] [topological_space M] [topological_add_group M] {v : vector_measure α M} {w : vector_measure α N} : - -v ⊥ᵥ w ↔ v ⊥ᵥ w := + -v ⟂ᵥ w ↔ v ⟂ᵥ w := ⟨λ h, neg_neg v ▸ h.neg_left, neg_left⟩ @[simp] lemma neg_right_iff {N : Type*} [add_comm_group N] [topological_space N] [topological_add_group N] {v : vector_measure α M} {w : vector_measure α N} : - v ⊥ᵥ -w ↔ v ⊥ᵥ w := + v ⟂ᵥ -w ↔ v ⟂ᵥ w := ⟨λ h, neg_neg w ▸ h.neg_right, neg_right⟩ end mutually_singular From 3592f45eeaf90143256b1dee3c20aade62924b0c Mon Sep 17 00:00:00 2001 From: negiizhao Date: Sun, 12 Feb 2023 10:53:06 +0000 Subject: [PATCH 0468/1291] chore(order/filter/filter_product): fix `const_lt`, make `const_lt_iff` `@[simp, norm_cast]` (#17442) --- src/order/filter/filter_product.lean | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/src/order/filter/filter_product.lean b/src/order/filter/filter_product.lean index 1aab2c5ab3f41..e2c823a2f90d4 100644 --- a/src/order/filter/filter_product.lean +++ b/src/order/filter/filter_product.lean @@ -47,8 +47,9 @@ by simp only [lt_iff_le_not_le, eventually_and, coe_le, eventually_not, eventual lemma coe_pos [preorder β] [has_zero β] {f : α → β} : 0 < (f : β*) ↔ ∀* x, 0 < f x := coe_lt -lemma const_lt [preorder β] {x y : β} : x ≤ y → (↑x : β*) ≤ ↑y := lift_rel_const +lemma const_lt [preorder β] {x y : β} : x < y → (↑x : β*) < ↑y := coe_lt.mpr ∘ lift_rel_const +@[simp, norm_cast] lemma const_lt_iff [preorder β] {x y : β} : (↑x : β*) < ↑y ↔ x < y := coe_lt.trans lift_rel_const_iff From 74e62ed9b580509201110ae830a9f0b3e9660dbe Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Sun, 12 Feb 2023 19:22:14 +0000 Subject: [PATCH 0469/1291] chore(analysis/convex/specific_functions): remove unused import (#18431) --- src/analysis/convex/specific_functions.lean | 1 - 1 file changed, 1 deletion(-) diff --git a/src/analysis/convex/specific_functions.lean b/src/analysis/convex/specific_functions.lean index a6a8429d91153..8a07f38210f91 100644 --- a/src/analysis/convex/specific_functions.lean +++ b/src/analysis/convex/specific_functions.lean @@ -3,7 +3,6 @@ Copyright (c) 2020 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov, Sébastien Gouëzel -/ -import analysis.calculus.mean_value import analysis.special_functions.pow_deriv import analysis.special_functions.sqrt From 6c5f73fd6f6cc83122788a80a27cdd54663609f4 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 13 Feb 2023 00:00:55 +0000 Subject: [PATCH 0470/1291] feat(algebra/big_operators/basic): `finset.sum` under `mod` (#18364) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit and `∏ a in s, f a = b ^ s.card` if `f a = b` for all `a` --- src/algebra/big_operators/basic.lean | 20 +++++++++++++++++++ src/algebra/big_operators/multiset/basic.lean | 12 +++++++++++ src/data/list/big_operators/basic.lean | 13 ++++++++++++ 3 files changed, 45 insertions(+) diff --git a/src/algebra/big_operators/basic.lean b/src/algebra/big_operators/basic.lean index aa968bc5c9fee..6ff8391e6c7db 100644 --- a/src/algebra/big_operators/basic.lean +++ b/src/algebra/big_operators/basic.lean @@ -1131,6 +1131,10 @@ end @[simp, to_additive] lemma prod_const (b : β) : (∏ x in s, b) = b ^ s.card := (congr_arg _ $ s.val.map_const b).trans $ multiset.prod_replicate s.card b +@[to_additive sum_eq_card_nsmul] lemma prod_eq_pow_card {b : β} (hf : ∀ a ∈ s, f a = b) : + ∏ a in s, f a = b ^ s.card := +(prod_congr rfl hf).trans $ prod_const _ + @[to_additive] lemma pow_eq_prod_const (b : β) : ∀ n, b ^ n = ∏ k in range n, b := by simp @@ -1570,6 +1574,22 @@ lemma prod_unique_nonempty {α β : Type*} [comm_monoid β] [unique α] (∏ x in s, f x) = f default := by rw [h.eq_singleton_default, finset.prod_singleton] +lemma sum_nat_mod (s : finset α) (n : ℕ) (f : α → ℕ) : + (∑ i in s, f i) % n = (∑ i in s, f i % n) % n := +(multiset.sum_nat_mod _ _).trans $ by rw [finset.sum, multiset.map_map] + +lemma prod_nat_mod (s : finset α) (n : ℕ) (f : α → ℕ) : + (∏ i in s, f i) % n = (∏ i in s, f i % n) % n := +(multiset.prod_nat_mod _ _).trans $ by rw [finset.prod, multiset.map_map] + +lemma sum_int_mod (s : finset α) (n : ℤ) (f : α → ℤ) : + (∑ i in s, f i) % n = (∑ i in s, f i % n) % n := +(multiset.sum_int_mod _ _).trans $ by rw [finset.sum, multiset.map_map] + +lemma prod_int_mod (s : finset α) (n : ℤ) (f : α → ℤ) : + (∏ i in s, f i) % n = (∏ i in s, f i % n) % n := +(multiset.prod_int_mod _ _).trans $ by rw [finset.prod, multiset.map_map] + end finset namespace fintype diff --git a/src/algebra/big_operators/multiset/basic.lean b/src/algebra/big_operators/multiset/basic.lean index b159e8d7fb906..6ca614001b23d 100644 --- a/src/algebra/big_operators/multiset/basic.lean +++ b/src/algebra/big_operators/multiset/basic.lean @@ -418,6 +418,18 @@ lemma abs_sum_le_sum_abs [linear_ordered_add_comm_group α] {s : multiset α} : abs s.sum ≤ (s.map abs).sum := le_sum_of_subadditive _ abs_zero abs_add s +lemma sum_nat_mod (s : multiset ℕ) (n : ℕ) : s.sum % n = (s.map (% n)).sum % n := +by induction s using multiset.induction; simp [nat.add_mod, *] + +lemma prod_nat_mod (s : multiset ℕ) (n : ℕ) : s.prod % n = (s.map (% n)).prod % n := +by induction s using multiset.induction; simp [nat.mul_mod, *] + +lemma sum_int_mod (s : multiset ℤ) (n : ℤ) : s.sum % n = (s.map (% n)).sum % n := +by induction s using multiset.induction; simp [int.add_mod, *] + +lemma prod_int_mod (s : multiset ℤ) (n : ℤ) : s.prod % n = (s.map (% n)).prod % n := +by induction s using multiset.induction; simp [int.mul_mod, *] + end multiset @[to_additive] diff --git a/src/data/list/big_operators/basic.lean b/src/data/list/big_operators/basic.lean index c93af87dad977..4e6925515baf0 100644 --- a/src/data/list/big_operators/basic.lean +++ b/src/data/list/big_operators/basic.lean @@ -3,6 +3,7 @@ Copyright (c) 2017 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Floris van Doorn, Sébastien Gouëzel, Alex J. Best -/ +import data.int.order.basic import data.list.forall2 /-! @@ -543,6 +544,18 @@ by rw [div_eq_mul_inv, alternating_prod_cons'] end alternating +lemma sum_nat_mod (l : list ℕ) (n : ℕ) : l.sum % n = (l.map (% n)).sum % n := +by induction l; simp [nat.add_mod, *] + +lemma prod_nat_mod (l : list ℕ) (n : ℕ) : l.prod % n = (l.map (% n)).prod % n := +by induction l; simp [nat.mul_mod, *] + +lemma sum_int_mod (l : list ℤ) (n : ℤ) : l.sum % n = (l.map (% n)).sum % n := +by induction l; simp [int.add_mod, *] + +lemma prod_int_mod (l : list ℤ) (n : ℤ) : l.prod % n = (l.map (% n)).prod % n := +by induction l; simp [int.mul_mod, *] + end list section monoid_hom From 4ea65cad5ad088877e021da13201d55a93c92366 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 13 Feb 2023 02:56:32 +0000 Subject: [PATCH 0471/1291] =?UTF-8?q?chore(algebra/algebra/basic):=20missi?= =?UTF-8?q?ng=20lemmas=20about=20the=20`=E2=86=91`=20spelling=20of=20`alge?= =?UTF-8?q?bra=5Fmap`=20(#18422)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This only really scratches the surface; in general we can't afford to have two different spellings here as it forces us to duplicate every result about `ring_hom`. A previously refactor (#17048) made these lemmas unecessary, but it was reverted to aid with the port. Since these lemmas are needed in mathlib3 PRs, lets have them anyway. For consistency with the existing `algebra_map.coe_mul` etc, these are not `simp` lemmas. Making them `simp` lemmas makes a bunch of downstream `simp` lemmas redundant; so perhaps worth exploring in a future PR. --- src/algebra/algebra/basic.lean | 30 ++++++++++++++++++++++++------ 1 file changed, 24 insertions(+), 6 deletions(-) diff --git a/src/algebra/algebra/basic.lean b/src/algebra/algebra/basic.lean index 9ce4bffda4eca..827e2cf0bc4b5 100644 --- a/src/algebra/algebra/basic.lean +++ b/src/algebra/algebra/basic.lean @@ -170,20 +170,38 @@ attribute [to_additive] coe_prod end comm_semiring_comm_semiring section field_nontrivial - variables {R A : Type*} [field R] [comm_semiring A] [nontrivial A] [algebra R A] @[norm_cast, simp] lemma coe_inj {a b : R} : (↑a : A) = ↑b ↔ a = b := -⟨λ h, (algebra_map R A).injective h, by rintro rfl; refl⟩ +(algebra_map R A).injective.eq_iff @[norm_cast, simp] lemma lift_map_eq_zero_iff (a : R) : (↑a : A) = 0 ↔ a = 0 := -begin - rw (show (0 : A) = ↑(0 : R), from (map_zero (algebra_map R A)).symm), - norm_cast, -end +map_eq_zero_iff _ (algebra_map R A).injective end field_nontrivial +section semifield_semidivision_ring +variables {R : Type*} (A : Type*) [semifield R] [division_semiring A] [algebra R A] + +@[norm_cast] lemma coe_inv (r : R) : ↑(r⁻¹) = ((↑r)⁻¹ : A) := +map_inv₀ (algebra_map R A) r + +@[norm_cast] lemma coe_div (r s : R) : ↑(r / s) = (↑r / ↑s : A) := +map_div₀ (algebra_map R A) r s + +@[norm_cast] lemma coe_zpow (r : R) (z : ℤ) : ↑(r ^ z) = (↑r ^ z : A) := +map_zpow₀ (algebra_map R A) r z + +end semifield_semidivision_ring + +section field_division_ring +variables (R A : Type*) [field R] [division_ring A] [algebra R A] + +@[norm_cast] lemma coe_rat_cast (q : ℚ) : ↑(q : R) = (q : A) := +map_rat_cast (algebra_map R A) q + +end field_division_ring + end algebra_map /-- Creating an algebra from a morphism to the center of a semiring. -/ From 34ee86e6a59d911a8e4f89b68793ee7577ae79c7 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Mon, 13 Feb 2023 07:22:29 +0000 Subject: [PATCH 0472/1291] chore(*): add mathlib4 synchronization comments (#18430) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.order.module` * `algebra.order.nonneg.field` * `algebra.order.nonneg.ring` * `category_theory.bicategory.strict` * `category_theory.eq_to_hom` * `category_theory.equivalence` * `category_theory.functor.const` * `category_theory.opposites` * `data.analysis.filter` * `data.dfinsupp.lex` * `data.fin.tuple.nat_antidiagonal` * `data.finite.card` * `data.finset.mul_antidiagonal` * `data.list.indexes` * `data.rat.denumerable` * `data.real.pointwise` * `group_theory.abelianization` * `group_theory.commutator` * `group_theory.finiteness` * `group_theory.group_action.quotient` * `measure_theory.measurable_space` * `measure_theory.pi_system` * `ring_theory.subring.basic` * `set_theory.cardinal.basic` * `set_theory.cardinal.finite` * `topology.algebra.semigroup` * `topology.compact_open` * `topology.continuous_function.t0_sierpinski` * `topology.discrete_quotient` * `topology.noetherian_space` * `topology.order.lower_topology` * `topology.partial` * `topology.sets.closeds` * `topology.sets.opens` * `topology.spectral.hom` * `topology.uniform_space.abstract_completion` * `topology.uniform_space.equiv` --- src/algebra/order/module.lean | 3 +++ src/algebra/order/nonneg/field.lean | 3 +++ src/algebra/order/nonneg/ring.lean | 3 +++ src/category_theory/bicategory/strict.lean | 3 +++ src/category_theory/eq_to_hom.lean | 3 +++ src/category_theory/equivalence.lean | 3 +++ src/category_theory/functor/const.lean | 3 +++ src/category_theory/opposites.lean | 3 +++ src/data/analysis/filter.lean | 3 +++ src/data/dfinsupp/lex.lean | 3 +++ src/data/fin/tuple/nat_antidiagonal.lean | 3 +++ src/data/finite/card.lean | 3 +++ src/data/finset/mul_antidiagonal.lean | 3 +++ src/data/list/indexes.lean | 3 +++ src/data/rat/denumerable.lean | 3 +++ src/data/real/pointwise.lean | 3 +++ src/group_theory/abelianization.lean | 3 +++ src/group_theory/commutator.lean | 3 +++ src/group_theory/finiteness.lean | 3 +++ src/group_theory/group_action/quotient.lean | 3 +++ src/measure_theory/measurable_space.lean | 3 +++ src/measure_theory/pi_system.lean | 3 +++ src/ring_theory/subring/basic.lean | 3 +++ src/set_theory/cardinal/basic.lean | 3 +++ src/set_theory/cardinal/finite.lean | 3 +++ src/topology/algebra/semigroup.lean | 3 +++ src/topology/compact_open.lean | 3 +++ src/topology/continuous_function/t0_sierpinski.lean | 3 +++ src/topology/discrete_quotient.lean | 3 +++ src/topology/noetherian_space.lean | 3 +++ src/topology/order/lower_topology.lean | 3 +++ src/topology/partial.lean | 3 +++ src/topology/sets/closeds.lean | 3 +++ src/topology/sets/opens.lean | 3 +++ src/topology/spectral/hom.lean | 3 +++ src/topology/uniform_space/abstract_completion.lean | 3 +++ src/topology/uniform_space/equiv.lean | 3 +++ 37 files changed, 111 insertions(+) diff --git a/src/algebra/order/module.lean b/src/algebra/order/module.lean index 8c3ef26904dd8..f4b75426245d9 100644 --- a/src/algebra/order/module.lean +++ b/src/algebra/order/module.lean @@ -8,6 +8,9 @@ import algebra.order.smul /-! # Ordered module +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we provide lemmas about `ordered_smul` that hold once a module structure is present. ## References diff --git a/src/algebra/order/nonneg/field.lean b/src/algebra/order/nonneg/field.lean index f3f9b3349aa6a..916655f8be1f5 100644 --- a/src/algebra/order/nonneg/field.lean +++ b/src/algebra/order/nonneg/field.lean @@ -11,6 +11,9 @@ import algebra.order.field.canonical.defs /-! # Semifield structure on the type of nonnegative elements +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines instances and prove some properties about the nonnegative elements `{x : α // 0 ≤ x}` of an arbitrary type `α`. diff --git a/src/algebra/order/nonneg/ring.lean b/src/algebra/order/nonneg/ring.lean index ea6fc73374301..3cd3e4242239c 100644 --- a/src/algebra/order/nonneg/ring.lean +++ b/src/algebra/order/nonneg/ring.lean @@ -13,6 +13,9 @@ import order.lattice_intervals /-! # The type of nonnegative elements +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines instances and prove some properties about the nonnegative elements `{x : α // 0 ≤ x}` of an arbitrary type `α`. diff --git a/src/category_theory/bicategory/strict.lean b/src/category_theory/bicategory/strict.lean index 7cd3b4d8c3f89..692dd0374e06c 100644 --- a/src/category_theory/bicategory/strict.lean +++ b/src/category_theory/bicategory/strict.lean @@ -9,6 +9,9 @@ import category_theory.bicategory.basic /-! # Strict bicategories +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A bicategory is called `strict` if the left unitors, the right unitors, and the associators are isomorphisms given by equalities. diff --git a/src/category_theory/eq_to_hom.lean b/src/category_theory/eq_to_hom.lean index 4bdeaed440d40..0ae887d5de674 100644 --- a/src/category_theory/eq_to_hom.lean +++ b/src/category_theory/eq_to_hom.lean @@ -8,6 +8,9 @@ import category_theory.opposites /-! # Morphisms from equations between objects. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + When working categorically, sometimes one encounters an equation `h : X = Y` between objects. Your initial aversion to this is natural and appropriate: diff --git a/src/category_theory/equivalence.lean b/src/category_theory/equivalence.lean index 49fa891c5602a..f421667168bf9 100644 --- a/src/category_theory/equivalence.lean +++ b/src/category_theory/equivalence.lean @@ -12,6 +12,9 @@ import tactic.slice /-! # Equivalence of categories +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + An equivalence of categories `C` and `D` is a pair of functors `F : C ⥤ D` and `G : D ⥤ C` such that `η : 𝟭 C ≅ F ⋙ G` and `ε : G ⋙ F ≅ 𝟭 D`. In many situations, equivalences are a better notion of "sameness" of categories than the stricter isomorphims of categories. diff --git a/src/category_theory/functor/const.lean b/src/category_theory/functor/const.lean index e1b08fce78845..ae1fc08e30df6 100644 --- a/src/category_theory/functor/const.lean +++ b/src/category_theory/functor/const.lean @@ -8,6 +8,9 @@ import category_theory.opposites /-! # The constant functor +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + `const J : C ⥤ (J ⥤ C)` is the functor that sends an object `X : C` to the functor `J ⥤ C` sending every object in `J` to `X`, and every morphism to `𝟙 X`. diff --git a/src/category_theory/opposites.lean b/src/category_theory/opposites.lean index a5917dbb5572f..b65f710780460 100644 --- a/src/category_theory/opposites.lean +++ b/src/category_theory/opposites.lean @@ -8,6 +8,9 @@ import category_theory.equivalence /-! # Opposite categories +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We provide a category instance on `Cᵒᵖ`. The morphisms `X ⟶ Y` are defined to be the morphisms `unop Y ⟶ unop X` in `C`. diff --git a/src/data/analysis/filter.lean b/src/data/analysis/filter.lean index 85a8a7fa08cf7..edff2cd7db0d5 100644 --- a/src/data/analysis/filter.lean +++ b/src/data/analysis/filter.lean @@ -8,6 +8,9 @@ import order.filter.cofinite /-! # Computational realization of filters (experimental) +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file provides infrastructure to compute with filters. ## Main declarations diff --git a/src/data/dfinsupp/lex.lean b/src/data/dfinsupp/lex.lean index 91aa1fa74b091..3f52398c4be95 100644 --- a/src/data/dfinsupp/lex.lean +++ b/src/data/dfinsupp/lex.lean @@ -10,6 +10,9 @@ import order.well_founded_set /-! # Lexicographic order on finitely supported dependent functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the lexicographic order on `dfinsupp`. -/ diff --git a/src/data/fin/tuple/nat_antidiagonal.lean b/src/data/fin/tuple/nat_antidiagonal.lean index a6cf499a743a3..e6aa9a68a4b4b 100644 --- a/src/data/fin/tuple/nat_antidiagonal.lean +++ b/src/data/fin/tuple/nat_antidiagonal.lean @@ -11,6 +11,9 @@ import logic.equiv.fin /-! # Collections of tuples of naturals with the same sum +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file generalizes `list.nat.antidiagonal n`, `multiset.nat.antidiagonal n`, and `finset.nat.antidiagonal n` from the pair of elements `x : ℕ × ℕ` such that `n = x.1 + x.2`, to the sequence of elements `x : fin k → ℕ` such that `n = ∑ i, x i`. diff --git a/src/data/finite/card.lean b/src/data/finite/card.lean index e916fa3a01927..4191dc9366418 100644 --- a/src/data/finite/card.lean +++ b/src/data/finite/card.lean @@ -9,6 +9,9 @@ import set_theory.cardinal.finite # Cardinality of finite types +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The cardinality of a finite type `α` is given by `nat.card α`. This function has the "junk value" of `0` for infinite types, but to ensure the function has valid output, one just needs to know that it's possible to produce a `finite` instance diff --git a/src/data/finset/mul_antidiagonal.lean b/src/data/finset/mul_antidiagonal.lean index 701ae27716717..a03bf9fff06a8 100644 --- a/src/data/finset/mul_antidiagonal.lean +++ b/src/data/finset/mul_antidiagonal.lean @@ -8,6 +8,9 @@ import data.set.mul_antidiagonal /-! # Multiplication antidiagonal as a `finset`. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We construct the `finset` of all pairs of an element in `s` and an element in `t` that multiply to `a`, given that `s` and `t` are well-ordered.-/ diff --git a/src/data/list/indexes.lean b/src/data/list/indexes.lean index 99fda0119a68e..f8264c39876d8 100644 --- a/src/data/list/indexes.lean +++ b/src/data/list/indexes.lean @@ -9,6 +9,9 @@ import data.list.range /-! # Lemmas about list.*_with_index functions. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Some specification lemmas for `list.map_with_index`, `list.mmap_with_index`, `list.foldl_with_index` and `list.foldr_with_index`. -/ diff --git a/src/data/rat/denumerable.lean b/src/data/rat/denumerable.lean index 37120717f6386..6c0d6a97d7e8c 100644 --- a/src/data/rat/denumerable.lean +++ b/src/data/rat/denumerable.lean @@ -8,6 +8,9 @@ import set_theory.cardinal.basic /-! # Denumerability of ℚ +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves that ℚ is infinite, denumerable, and deduces that it has cardinality `omega`. -/ diff --git a/src/data/real/pointwise.lean b/src/data/real/pointwise.lean index f3b74ce52f747..98e07256f07b9 100644 --- a/src/data/real/pointwise.lean +++ b/src/data/real/pointwise.lean @@ -9,6 +9,9 @@ import data.real.basic /-! # Pointwise operations on sets of reals +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file relates `Inf (a • s)`/`Sup (a • s)` with `a • Inf s`/`a • Sup s` for `s : set ℝ`. From these, it relates `⨅ i, a • f i` / `⨆ i, a • f i` with `a • (⨅ i, f i)` / `a • (⨆ i, f i)`, diff --git a/src/group_theory/abelianization.lean b/src/group_theory/abelianization.lean index bf32cb1c68bc0..6c9fdf44175e5 100644 --- a/src/group_theory/abelianization.lean +++ b/src/group_theory/abelianization.lean @@ -10,6 +10,9 @@ import group_theory.finiteness /-! # The abelianization of a group +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the commutator and the abelianization of a group. It furthermore prepares for the result that the abelianization is left adjoint to the forgetful functor from abelian groups to groups, which can be found in `algebra/category/Group/adjunctions`. diff --git a/src/group_theory/commutator.lean b/src/group_theory/commutator.lean index 3206dab1db37a..8d2ca4052460a 100644 --- a/src/group_theory/commutator.lean +++ b/src/group_theory/commutator.lean @@ -11,6 +11,9 @@ import tactic.group /-! # Commutators of Subgroups +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + If `G` is a group and `H₁ H₂ : subgroup G` then the commutator `⁅H₁, H₂⁆ : subgroup G` is the subgroup of `G` generated by the commutators `h₁ * h₂ * h₁⁻¹ * h₂⁻¹`. diff --git a/src/group_theory/finiteness.lean b/src/group_theory/finiteness.lean index a8a109cd90b0e..7a47344442139 100644 --- a/src/group_theory/finiteness.lean +++ b/src/group_theory/finiteness.lean @@ -14,6 +14,9 @@ import data.finset.preimage /-! # Finitely generated monoids and groups +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define finitely generated monoids and groups. See also `submodule.fg` and `module.finite` for finitely-generated modules. diff --git a/src/group_theory/group_action/quotient.lean b/src/group_theory/group_action/quotient.lean index a0c81928181a7..0c2d4a21eb7af 100644 --- a/src/group_theory/group_action/quotient.lean +++ b/src/group_theory/group_action/quotient.lean @@ -13,6 +13,9 @@ import group_theory.coset /-! # Properties of group actions involving quotient groups +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves properties of group actions which use the quotient group construction, notably * the orbit-stabilizer theorem `card_orbit_mul_card_stabilizer_eq_card_group` * the class formula `card_eq_sum_card_group_div_card_stabilizer'` diff --git a/src/measure_theory/measurable_space.lean b/src/measure_theory/measurable_space.lean index b240ba194da85..d9c69d5c5e234 100644 --- a/src/measure_theory/measurable_space.lean +++ b/src/measure_theory/measurable_space.lean @@ -15,6 +15,9 @@ import measure_theory.tactic /-! # Measurable spaces and measurable functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file provides properties of measurable spaces and the functions and isomorphisms between them. The definition of a measurable space is in `measure_theory.measurable_space_def`. diff --git a/src/measure_theory/pi_system.lean b/src/measure_theory/pi_system.lean index eeb339d764f33..6b51be17b3438 100644 --- a/src/measure_theory/pi_system.lean +++ b/src/measure_theory/pi_system.lean @@ -9,6 +9,9 @@ import measure_theory.measurable_space_def /-! # Induction principles for measurable sets, related to π-systems and λ-systems. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main statements * The main theorem of this file is Dynkin's π-λ theorem, which appears diff --git a/src/ring_theory/subring/basic.lean b/src/ring_theory/subring/basic.lean index de604180ee48e..68334a8fdd49d 100644 --- a/src/ring_theory/subring/basic.lean +++ b/src/ring_theory/subring/basic.lean @@ -10,6 +10,9 @@ import ring_theory.subsemiring.basic /-! # Subrings +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Let `R` be a ring. This file defines the "bundled" subring type `subring R`, a type whose terms correspond to subrings of `R`. This is the preferred way to talk about subrings in mathlib. Unbundled subrings (`s : set R` and `is_subring s`) diff --git a/src/set_theory/cardinal/basic.lean b/src/set_theory/cardinal/basic.lean index 880b32b1fcaf4..e64e91c84d436 100644 --- a/src/set_theory/cardinal/basic.lean +++ b/src/set_theory/cardinal/basic.lean @@ -16,6 +16,9 @@ import tactic.positivity /-! # Cardinal Numbers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define cardinal numbers as a quotient of types under the equivalence relation of equinumerity. ## Main definitions diff --git a/src/set_theory/cardinal/finite.lean b/src/set_theory/cardinal/finite.lean index 537d6afb6d7e9..5495b095f9091 100644 --- a/src/set_theory/cardinal/finite.lean +++ b/src/set_theory/cardinal/finite.lean @@ -9,6 +9,9 @@ import set_theory.cardinal.basic /-! # Finite Cardinality Functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main Definitions * `nat.card α` is the cardinality of `α` as a natural number. diff --git a/src/topology/algebra/semigroup.lean b/src/topology/algebra/semigroup.lean index 0fed5ede93dee..c5ebed196f392 100644 --- a/src/topology/algebra/semigroup.lean +++ b/src/topology/algebra/semigroup.lean @@ -8,6 +8,9 @@ import topology.separation /-! # Idempotents in topological semigroups +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file provides a sufficient condition for a semigroup `M` to contain an idempotent (i.e. an element `m` such that `m * m = m `), namely that `M` is a nonempty compact Hausdorff space where right-multiplication by constants is continuous. diff --git a/src/topology/compact_open.lean b/src/topology/compact_open.lean index a1fb25e06a393..78fd62cddd93d 100644 --- a/src/topology/compact_open.lean +++ b/src/topology/compact_open.lean @@ -12,6 +12,9 @@ import topology.maps /-! # The compact-open topology +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we define the compact-open topology on the set of continuous maps between two topological spaces. diff --git a/src/topology/continuous_function/t0_sierpinski.lean b/src/topology/continuous_function/t0_sierpinski.lean index 0c511fb9044e3..5c8bf236ec146 100644 --- a/src/topology/continuous_function/t0_sierpinski.lean +++ b/src/topology/continuous_function/t0_sierpinski.lean @@ -10,6 +10,9 @@ import topology.continuous_function.basic /-! # Any T0 space embeds in a product of copies of the Sierpinski space. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We consider `Prop` with the Sierpinski topology. If `X` is a topological space, there is a continuous map `product_of_mem_opens` from `X` to `opens X → Prop` which is the product of the maps `X → Prop` given by `x ↦ x ∈ u`. diff --git a/src/topology/discrete_quotient.lean b/src/topology/discrete_quotient.lean index 801bf98c68845..e1b31901b9b35 100644 --- a/src/topology/discrete_quotient.lean +++ b/src/topology/discrete_quotient.lean @@ -11,6 +11,9 @@ import topology.locally_constant.basic # Discrete quotients of a topological space. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the type of discrete quotients of a topological space, denoted `discrete_quotient X`. To avoid quantifying over types, we model such quotients as setoids whose equivalence classes are clopen. diff --git a/src/topology/noetherian_space.lean b/src/topology/noetherian_space.lean index 2a1f3fda92787..18d85a615a0f1 100644 --- a/src/topology/noetherian_space.lean +++ b/src/topology/noetherian_space.lean @@ -9,6 +9,9 @@ import topology.sets.closeds /-! # Noetherian space +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A Noetherian space is a topological space that satisfies any of the following equivalent conditions: - `well_founded ((>) : opens α → opens α → Prop)` - `well_founded ((<) : closeds α → closeds α → Prop)` diff --git a/src/topology/order/lower_topology.lean b/src/topology/order/lower_topology.lean index f5007759beb21..0ed69357c94a1 100644 --- a/src/topology/order/lower_topology.lean +++ b/src/topology/order/lower_topology.lean @@ -10,6 +10,9 @@ import order.hom.complete_lattice /-! # Lower topology +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file introduces the lower topology on a preorder as the topology generated by the complements of the closed intervals to infinity. diff --git a/src/topology/partial.lean b/src/topology/partial.lean index 4ebcbae24f64e..82cbe064b1c9c 100644 --- a/src/topology/partial.lean +++ b/src/topology/partial.lean @@ -9,6 +9,9 @@ import order.filter.partial /-! # Partial functions and topological spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove properties of `filter.ptendsto` etc in topological spaces. We also introduce `pcontinuous`, a version of `continuous` for partially defined functions. -/ diff --git a/src/topology/sets/closeds.lean b/src/topology/sets/closeds.lean index 9cc9db98e3f94..dbe085444079e 100644 --- a/src/topology/sets/closeds.lean +++ b/src/topology/sets/closeds.lean @@ -8,6 +8,9 @@ import topology.sets.opens /-! # Closed sets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define a few types of closed sets in a topological space. ## Main Definitions diff --git a/src/topology/sets/opens.lean b/src/topology/sets/opens.lean index a64f094a3f5b2..5b13c8ea4c6fb 100644 --- a/src/topology/sets/opens.lean +++ b/src/topology/sets/opens.lean @@ -13,6 +13,9 @@ import tactic.auto_cases /-! # Open sets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Summary We define the subtype of open sets in a topological space. diff --git a/src/topology/spectral/hom.lean b/src/topology/spectral/hom.lean index 197958115ca51..45962b21f8052 100644 --- a/src/topology/spectral/hom.lean +++ b/src/topology/spectral/hom.lean @@ -8,6 +8,9 @@ import topology.continuous_function.basic /-! # Spectral maps +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines spectral maps. A map is spectral when it's continuous and the preimage of a compact open set is compact open. diff --git a/src/topology/uniform_space/abstract_completion.lean b/src/topology/uniform_space/abstract_completion.lean index f948d01536302..c419423a9b3db 100644 --- a/src/topology/uniform_space/abstract_completion.lean +++ b/src/topology/uniform_space/abstract_completion.lean @@ -9,6 +9,9 @@ import topology.uniform_space.equiv /-! # Abstract theory of Hausdorff completions of uniform spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file characterizes Hausdorff completions of a uniform space α as complete Hausdorff spaces equipped with a map from α which has dense image and induce the original uniform structure on α. Assuming these properties we "extend" uniformly continuous maps from α to complete Hausdorff spaces diff --git a/src/topology/uniform_space/equiv.lean b/src/topology/uniform_space/equiv.lean index 5d7a81942d349..7103f02e8ca90 100644 --- a/src/topology/uniform_space/equiv.lean +++ b/src/topology/uniform_space/equiv.lean @@ -11,6 +11,9 @@ import topology.uniform_space.pi /-! # Uniform isomorphisms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines uniform isomorphisms between two uniform spaces. They are bijections with both directions uniformly continuous. We denote uniform isomorphisms with the notation `≃ᵤ`. From 48085f140e684306f9e7da907cd5932056d1aded Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 13 Feb 2023 19:19:10 +0000 Subject: [PATCH 0473/1291] chore(algebra/quaternion): generalize the algebra instance (#18382) To prevent this generalization forming diamonds, we have to now populate the `nsmul`, `zsmul`, and `qsmul` fields. I use this in another PR to talk about the `R` algebra structure of `quaternion (dual_number R)`. Even without that motivation, this change means that the existing `smul` lemmas now apply to `nsmul`, `zsmul`, and `rat_smul` where previously they did not. --- src/algebra/quaternion.lean | 103 +++++++++++++++++++++++++---------- src/analysis/quaternion.lean | 2 +- 2 files changed, 74 insertions(+), 31 deletions(-) diff --git a/src/algebra/quaternion.lean b/src/algebra/quaternion.lean index 8822b8292c4e3..8ecb6bf2fb2c8 100644 --- a/src/algebra/quaternion.lean +++ b/src/algebra/quaternion.lean @@ -82,7 +82,7 @@ def equiv_tuple {R : Type*} (c₁ c₂ : R) : ℍ[R, c₁, c₂] ≃ (fin 4 → @[simp] lemma mk.eta {R : Type*} {c₁ c₂} : ∀ a : ℍ[R, c₁, c₂], mk a.1 a.2 a.3 a.4 = a | ⟨a₁, a₂, a₃, a₄⟩ := rfl -variables {R : Type*} [comm_ring R] {c₁ c₂ : R} (r x y z : R) (a b c : ℍ[R, c₁, c₂]) +variables {S T R : Type*} [comm_ring R] {c₁ c₂ : R} (r x y z : R) (a b c : ℍ[R, c₁, c₂]) instance : has_coe_t R (ℍ[R, c₁, c₂]) := ⟨λ x, ⟨x, 0, 0, 0⟩⟩ @@ -153,14 +153,46 @@ rfl a₁ * b₃ + c₁ * a₂ * b₄ + a₃ * b₁ - c₁ * a₄ * b₂, a₁ * b₄ + a₂ * b₃ - a₃ * b₂ + a₄ * b₁⟩ := rfl +section +variables [has_smul S R] [has_smul T R] (s : S) + +/- +The `ring R` argument is not used, but it's also much stronger than the other definitions in this +file need; for instance `quaternion_algebra.has_zero` only really needs `has_zero R`. For +simplicity we just keep things consistent. +-/ +@[nolint unused_arguments] +instance : has_smul S ℍ[R, c₁, c₂] := +{ smul := λ s a, ⟨s • a.1, s • a.2, s • a.3, s • a.4⟩ } + +instance [has_smul S T] [is_scalar_tower S T R] : is_scalar_tower S T ℍ[R, c₁, c₂] := +{ smul_assoc := λ s t x, by ext; exact smul_assoc _ _ _ } + +instance [smul_comm_class S T R] : smul_comm_class S T ℍ[R, c₁, c₂] := +{ smul_comm := λ s t x, by ext; exact smul_comm _ _ _ } + +@[simp] lemma smul_re : (s • a).re = s • a.re := rfl +@[simp] lemma smul_im_i : (s • a).im_i = s • a.im_i := rfl +@[simp] lemma smul_im_j : (s • a).im_j = s • a.im_j := rfl +@[simp] lemma smul_im_k : (s • a).im_k = s • a.im_k := rfl + +@[simp] lemma smul_mk (re im_i im_j im_k : R) : + s • (⟨re, im_i, im_j, im_k⟩ : ℍ[R, c₁, c₂]) = ⟨s • re, s • im_i, s • im_j, s • im_k⟩ := rfl + +end + +@[simp, norm_cast] lemma coe_smul [smul_zero_class S R] (s : S) (r : R) : + (↑(s • r) : ℍ[R, c₁, c₂]) = s • ↑r := +ext _ _ rfl (smul_zero s).symm (smul_zero s).symm (smul_zero s).symm + instance : add_comm_group ℍ[R, c₁, c₂] := by refine_struct { add := (+), neg := has_neg.neg, sub := has_sub.sub, zero := (0 : ℍ[R, c₁, c₂]), - zsmul := @zsmul_rec _ ⟨(0 : ℍ[R, c₁, c₂])⟩ ⟨(+)⟩ ⟨has_neg.neg⟩, - nsmul := @nsmul_rec _ ⟨(0 : ℍ[R, c₁, c₂])⟩ ⟨(+)⟩ }; + nsmul := (•), + zsmul := (•), }; intros; try { refl }; ext; simp; ring_exp instance : add_group_with_one ℍ[R, c₁, c₂] := @@ -196,23 +228,20 @@ by refine_struct .. quaternion_algebra.add_comm_group }; intros; try { refl }; ext; simp; ring_exp -instance : algebra R ℍ[R, c₁, c₂] := -{ smul := λ r a, ⟨r * a.1, r * a.2, r * a.3, r * a.4⟩, - to_fun := coe, - map_one' := rfl, - map_zero' := rfl, - map_mul' := λ x y, by ext; simp, - map_add' := λ x y, by ext; simp, - smul_def' := λ r x, by ext; simp, - commutes' := λ r x, by ext; simp [mul_comm] } - -@[simp] lemma smul_re : (r • a).re = r • a.re := rfl -@[simp] lemma smul_im_i : (r • a).im_i = r • a.im_i := rfl -@[simp] lemma smul_im_j : (r • a).im_j = r • a.im_j := rfl -@[simp] lemma smul_im_k : (r • a).im_k = r • a.im_k := rfl +@[norm_cast, simp] lemma coe_mul : ((x * y : R) : ℍ[R, c₁, c₂]) = x * y := +by ext; simp -@[simp] lemma smul_mk (re im_i im_j im_k : R) : - r • (⟨re, im_i, im_j, im_k⟩ : ℍ[R, c₁, c₂]) = ⟨r • re, r • im_i, r • im_j, r • im_k⟩ := rfl +-- TODO: add weaker `mul_action`, `distrib_mul_action`, and `module` instances (and repeat them +-- for `ℍ[R]`) +instance [comm_semiring S] [algebra S R] : algebra S ℍ[R, c₁, c₂] := +{ smul := (•), + to_fun := λ s, coe (algebra_map S R s), + map_one' := by simpa only [map_one], + map_zero' := by simpa only [map_zero], + map_mul' := λ x y, by rw [map_mul, coe_mul], + map_add' := λ x y, by rw [map_add, coe_add], + smul_def' := λ s x, by ext; simp [algebra.smul_def], + commutes' := λ s x, by ext; simp [algebra.commutes] } lemma algebra_map_eq (r : R) : algebra_map R ℍ[R,c₁,c₂] r = ⟨r, 0, 0, 0⟩ := rfl @@ -271,9 +300,6 @@ end @[norm_cast, simp] lemma coe_sub : ((x - y : R) : ℍ[R, c₁, c₂]) = x - y := (algebra_map R ℍ[R, c₁, c₂]).map_sub x y -@[norm_cast, simp] lemma coe_mul : ((x * y : R) : ℍ[R, c₁, c₂]) = x * y := -(algebra_map R ℍ[R, c₁, c₂]).map_mul x y - @[norm_cast, simp] lemma coe_pow (n : ℕ) : (↑(x ^ n) : ℍ[R, c₁, c₂]) = ↑x ^ n := (algebra_map R ℍ[R, c₁, c₂]).map_pow x n @@ -351,7 +377,9 @@ by rw [←coe_nat_cast, conj_coe] @[simp, norm_cast] lemma conj_int_cast (z : ℤ) : conj (z : ℍ[R, c₁, c₂]) = z := by rw [←coe_int_cast, conj_coe] -lemma conj_smul : conj (r • a) = r • conj a := conj.map_smul r a +@[simp] lemma conj_smul [monoid S] [distrib_mul_action S R] (s : S) (a : ℍ[R, c₁, c₂]) : + conj (s • a) = s • conj a := +ext _ _ rfl (smul_neg _ _).symm (smul_neg _ _).symm (smul_neg _ _).symm @[simp] lemma conj_one : conj (1 : ℍ[R, c₁, c₂]) = 1 := conj_coe 1 @@ -426,14 +454,19 @@ quaternion_algebra.equiv_tuple _ _ namespace quaternion -variables {R : Type*} [comm_ring R] (r x y z : R) (a b c : ℍ[R]) +variables {S T R : Type*} [comm_ring R] (r x y z : R) (a b c : ℍ[R]) export quaternion_algebra (re im_i im_j im_k) instance : has_coe_t R ℍ[R] := quaternion_algebra.has_coe_t instance : ring ℍ[R] := quaternion_algebra.ring instance : inhabited ℍ[R] := quaternion_algebra.inhabited -instance : algebra R ℍ[R] := quaternion_algebra.algebra +instance [has_smul S R] : has_smul S ℍ[R] := quaternion_algebra.has_smul +instance [has_smul S T] [has_smul S R] [has_smul T R] [is_scalar_tower S T R] : + is_scalar_tower S T ℍ[R] := quaternion_algebra.is_scalar_tower +instance [has_smul S R] [has_smul T R] [smul_comm_class S T R] : + smul_comm_class S T ℍ[R] := quaternion_algebra.smul_comm_class +instance [comm_semiring S] [algebra S R] : algebra S ℍ[R] := quaternion_algebra.algebra instance : star_ring ℍ[R] := quaternion_algebra.star_ring @[ext] lemma ext : a.re = b.re → a.im_i = b.im_i → a.im_j = b.im_j → a.im_k = b.im_k → a = b := @@ -519,10 +552,14 @@ lemma coe_injective : function.injective (coe : R → ℍ[R]) := quaternion_alge @[simp] lemma coe_inj {x y : R} : (x : ℍ[R]) = y ↔ x = y := coe_injective.eq_iff -@[simp] lemma smul_re : (r • a).re = r • a.re := rfl -@[simp] lemma smul_im_i : (r • a).im_i = r • a.im_i := rfl -@[simp] lemma smul_im_j : (r • a).im_j = r • a.im_j := rfl -@[simp] lemma smul_im_k : (r • a).im_k = r • a.im_k := rfl +@[simp] lemma smul_re [has_smul S R] (s : S) : (s • a).re = s • a.re := rfl +@[simp] lemma smul_im_i [has_smul S R] (s : S) : (s • a).im_i = s • a.im_i := rfl +@[simp] lemma smul_im_j [has_smul S R] (s : S) : (s • a).im_j = s • a.im_j := rfl +@[simp] lemma smul_im_k [has_smul S R] (s : S) : (s • a).im_k = s • a.im_k := rfl + +@[simp, norm_cast] lemma coe_smul [smul_zero_class S R] (s : S) (r : R) : + (↑(s • r) : ℍ[R]) = s • ↑r := +quaternion_algebra.coe_smul _ _ lemma coe_commutes : ↑r * a = a * r := quaternion_algebra.coe_commutes r a @@ -589,7 +626,8 @@ quaternion_algebra.conj_nat_cast _ @[simp, norm_cast] lemma conj_int_cast (z : ℤ) : conj (z : ℍ[R]) = z := quaternion_algebra.conj_int_cast _ -@[simp] lemma conj_smul : conj (r • a) = r • conj a := a.conj_smul r +@[simp] lemma conj_smul [monoid S] [distrib_mul_action S R] (s : S) (a : ℍ[R]) : + conj (s • a) = s • conj a := quaternion_algebra.conj_smul _ _ @[simp] lemma conj_one : conj (1 : ℍ[R]) = 1 := conj_coe 1 @@ -724,6 +762,11 @@ map_zpow₀ (algebra_map R ℍ[R]) x z instance : division_ring ℍ[R] := { rat_cast := λ q, ↑(q : R), rat_cast_mk := λ n d hd h, by rw [rat.cast_mk', coe_mul, coe_int_cast, coe_inv, coe_nat_cast], + qsmul := (•), + qsmul_eq_mul' := λ q x, begin + rw coe_mul_eq_smul, + ext; exact division_ring.qsmul_eq_mul' _ _, + end, .. quaternion.group_with_zero, .. quaternion.ring } diff --git a/src/analysis/quaternion.lean b/src/analysis/quaternion.lean index b5c499aa8dd2a..d1c6d0ed97037 100644 --- a/src/analysis/quaternion.lean +++ b/src/analysis/quaternion.lean @@ -75,7 +75,7 @@ noncomputable instance : normed_division_ring ℍ := instance : normed_algebra ℝ ℍ := { norm_smul_le := λ a x, (norm_smul a x).le, - to_algebra := quaternion.algebra } + to_algebra := (quaternion.algebra : algebra ℝ ℍ) } instance : cstar_ring ℍ := { norm_star_mul_self := λ x, (norm_mul _ _).trans $ congr_arg (* ‖x‖) (norm_conj x) } From 26c71658340519485e2356d20cdca619e1cd4e19 Mon Sep 17 00:00:00 2001 From: Thomas Browning Date: Mon, 13 Feb 2023 22:42:15 +0000 Subject: [PATCH 0474/1291] refactor(combinatorics/configuration): Refactor `projective_plane` to extend `has_points` and `has_lines` (#18436) This PR refactors `projective_plane` to extend `has_points` and `has_lines`. For this to work in Lean3, we need to use `old_structure_cmd`. Co-authored-by: Eric Wieser --- src/combinatorics/configuration.lean | 43 ++++++++++++---------------- 1 file changed, 18 insertions(+), 25 deletions(-) diff --git a/src/combinatorics/configuration.lean b/src/combinatorics/configuration.lean index 36221045516c0..401a530781ab2 100644 --- a/src/combinatorics/configuration.lean +++ b/src/combinatorics/configuration.lean @@ -34,6 +34,8 @@ Together, these four statements say that any two of the following properties imp open_locale big_operators +set_option old_structure_cmd true + namespace configuration variables (P L : Type*) [has_mem P L] @@ -70,7 +72,7 @@ class has_lines extends nondegenerate P L := (mk_line : Π {p₁ p₂ : P} (h : p₁ ≠ p₂), L) (mk_line_ax : ∀ {p₁ p₂ : P} (h : p₁ ≠ p₂), p₁ ∈ mk_line h ∧ p₂ ∈ mk_line h) -open nondegenerate has_points has_lines +open nondegenerate has_points (mk_point mk_point_ax) has_lines (mk_line mk_line_ax) instance [nondegenerate P L] : nondegenerate (dual L) (dual P) := { exists_point := @exists_line P L _ _, @@ -79,11 +81,13 @@ instance [nondegenerate P L] : nondegenerate (dual L) (dual P) := instance [has_points P L] : has_lines (dual L) (dual P) := { mk_line := @mk_point P L _ _, - mk_line_ax := λ _ _, mk_point_ax } + mk_line_ax := λ _ _, mk_point_ax, + .. dual.nondegenerate _ _ } instance [has_lines P L] : has_points (dual L) (dual P) := { mk_point := @mk_line P L _ _, - mk_point_ax := λ _ _, mk_line_ax } + mk_point_ax := λ _ _, mk_line_ax, + .. dual.nondegenerate _ _ } lemma has_points.exists_unique_point [has_points P L] (l₁ l₂ : L) (hl : l₁ ≠ l₂) : ∃! p, p ∈ l₁ ∧ p ∈ l₂ := @@ -292,7 +296,8 @@ let this : ∀ l₁ l₂ : L, l₁ ≠ l₂ → ∃ p : P, p ∈ l₁ ∧ p ∈ exact ⟨q, (congr_arg _ (subtype.ext_iff.mp hq)).mp (mk_line_ax (this q)).2, q.2⟩, end in { mk_point := λ l₁ l₂ hl, classical.some (this l₁ l₂ hl), - mk_point_ax := λ l₁ l₂ hl, classical.some_spec (this l₁ l₂ hl) } + mk_point_ax := λ l₁ l₂ hl, classical.some_spec (this l₁ l₂ hl), + .. ‹has_lines P L› } /-- If a nondegenerate configuration has a unique point on any two lines, and if `|P| = |L|`, then there is a unique line through any two points. -/ @@ -300,41 +305,29 @@ noncomputable def has_points.has_lines [has_points P L] [fintype P] [fintype L] (h : fintype.card P = fintype.card L) : has_lines P L := let this := @has_lines.has_points (dual L) (dual P) _ _ _ _ h.symm in { mk_line := λ _ _, this.mk_point, - mk_line_ax := λ _ _, this.mk_point_ax } + mk_line_ax := λ _ _, this.mk_point_ax, + .. ‹has_points P L› } variables (P L) /-- A projective plane is a nondegenerate configuration in which every pair of lines has an intersection point, every pair of points has a line through them, and which has three points in general position. -/ -class projective_plane extends nondegenerate P L := -(mk_point : Π {l₁ l₂ : L} (h : l₁ ≠ l₂), P) -(mk_point_ax : ∀ {l₁ l₂ : L} (h : l₁ ≠ l₂), mk_point h ∈ l₁ ∧ mk_point h ∈ l₂) -(mk_line : Π {p₁ p₂ : P} (h : p₁ ≠ p₂), L) -(mk_line_ax : ∀ {p₁ p₂ : P} (h : p₁ ≠ p₂), p₁ ∈ mk_line h ∧ p₂ ∈ mk_line h) +class projective_plane extends has_points P L, has_lines P L := (exists_config : ∃ (p₁ p₂ p₃ : P) (l₁ l₂ l₃ : L), p₁ ∉ l₂ ∧ p₁ ∉ l₃ ∧ p₂ ∉ l₁ ∧ p₂ ∈ l₂ ∧ p₂ ∈ l₃ ∧ p₃ ∉ l₁ ∧ p₃ ∈ l₂ ∧ p₃ ∉ l₃) namespace projective_plane -@[priority 100] -- see Note [lower instance priority] -instance has_points [h : projective_plane P L] : has_points P L := { .. h } - -@[priority 100] -- see Note [lower instance priority] -instance has_lines [h : projective_plane P L] : has_lines P L := { .. h } - variables [projective_plane P L] instance : projective_plane (dual L) (dual P) := -{ mk_line := @mk_point P L _ _, - mk_line_ax := λ _ _, mk_point_ax, - mk_point := @mk_line P L _ _, - mk_point_ax := λ _ _, mk_line_ax, - exists_config := by - { obtain ⟨p₁, p₂, p₃, l₁, l₂, l₃, h₁₂, h₁₃, h₂₁, h₂₂, h₂₃, h₃₁, h₃₂, h₃₃⟩ := - @exists_config P L _ _, - exact ⟨l₁, l₂, l₃, p₁, p₂, p₃, h₂₁, h₃₁, h₁₂, h₂₂, h₃₂, h₁₃, h₂₃, h₃₃⟩ }, - .. dual.nondegenerate P L } +{ exists_config := + let ⟨p₁, p₂, p₃, l₁, l₂, l₃, h₁₂, h₁₃, h₂₁, h₂₂, h₂₃, h₃₁, h₃₂, h₃₃⟩ := + @exists_config P L _ _ in + ⟨l₁, l₂, l₃, p₁, p₂, p₃, h₂₁, h₃₁, h₁₂, h₂₂, h₃₂, h₁₃, h₂₃, h₃₃⟩, + .. dual.has_points _ _, + .. dual.has_lines _ _ } /-- The order of a projective plane is one less than the number of lines through an arbitrary point. Equivalently, it is one less than the number of points on an arbitrary line. -/ From ac34df03f74e6f797efd6991df2e3b7f7d8d33e0 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Tue, 14 Feb 2023 06:56:45 +0000 Subject: [PATCH 0475/1291] chore(*): add mathlib4 synchronization comments (#18438) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.order.kleene` * `category_theory.functor.inv_isos` * `combinatorics.composition` * `linear_algebra.basic` * `set_theory.ordinal.basic` * `topology.alexandroff` * `topology.algebra.mul_action` * `topology.algebra.order.left_right_lim` * `topology.quasi_separated` * `topology.sets.order` --- src/algebra/order/kleene.lean | 3 +++ src/category_theory/functor/inv_isos.lean | 3 +++ src/combinatorics/composition.lean | 3 +++ src/linear_algebra/basic.lean | 3 +++ src/set_theory/ordinal/basic.lean | 3 +++ src/topology/alexandroff.lean | 3 +++ src/topology/algebra/mul_action.lean | 3 +++ src/topology/algebra/order/left_right_lim.lean | 3 +++ src/topology/quasi_separated.lean | 3 +++ src/topology/sets/order.lean | 3 +++ 10 files changed, 30 insertions(+) diff --git a/src/algebra/order/kleene.lean b/src/algebra/order/kleene.lean index 0a2487a126cdf..ed2ef7a143c8e 100644 --- a/src/algebra/order/kleene.lean +++ b/src/algebra/order/kleene.lean @@ -11,6 +11,9 @@ import order.hom.complete_lattice /-! # Kleene Algebras +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines idempotent semirings and Kleene algebras, which are used extensively in the theory of computation. diff --git a/src/category_theory/functor/inv_isos.lean b/src/category_theory/functor/inv_isos.lean index 107ea4d853b5a..0dc3e4d9e6f97 100644 --- a/src/category_theory/functor/inv_isos.lean +++ b/src/category_theory/functor/inv_isos.lean @@ -9,6 +9,9 @@ import category_theory.eq_to_hom /-! # Natural isomorphisms with composition with inverses +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Definition of useful natural isomorphisms involving inverses of functors. These definitions cannot go in `category_theory/equivalence` because they require `eq_to_hom`. -/ diff --git a/src/combinatorics/composition.lean b/src/combinatorics/composition.lean index 235638b9f8347..bb6e5c5feb86a 100644 --- a/src/combinatorics/composition.lean +++ b/src/combinatorics/composition.lean @@ -10,6 +10,9 @@ import algebra.big_operators.fin /-! # Compositions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A composition of a natural number `n` is a decomposition `n = i₀ + ... + i_{k-1}` of `n` into a sum of positive integers. Combinatorially, it corresponds to a decomposition of `{0, ..., n-1}` into non-empty blocks of consecutive integers, where the `iⱼ` are the lengths of the blocks. diff --git a/src/linear_algebra/basic.lean b/src/linear_algebra/basic.lean index dfc8322aa2fbc..4fe744590ba1e 100644 --- a/src/linear_algebra/basic.lean +++ b/src/linear_algebra/basic.lean @@ -14,6 +14,9 @@ import data.finsupp.basic /-! # Linear algebra +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the basics of linear algebra. It sets up the "categorical/lattice structure" of modules over a ring, submodules, and linear maps. diff --git a/src/set_theory/ordinal/basic.lean b/src/set_theory/ordinal/basic.lean index 6c36f67c2e29f..bfd69d2b246ef 100644 --- a/src/set_theory/ordinal/basic.lean +++ b/src/set_theory/ordinal/basic.lean @@ -10,6 +10,9 @@ import set_theory.cardinal.basic /-! # Ordinals +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Ordinals are defined as equivalences of well-ordered sets under order isomorphism. They are endowed with a total order, where an ordinal is smaller than another one if it embeds into it as an initial segment (or, equivalently, in any way). This total order is well founded. diff --git a/src/topology/alexandroff.lean b/src/topology/alexandroff.lean index 2cc85d25521bb..fc0fae833aaa8 100644 --- a/src/topology/alexandroff.lean +++ b/src/topology/alexandroff.lean @@ -10,6 +10,9 @@ import topology.sets.opens /-! # The Alexandroff Compactification +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We construct the Alexandroff compactification (the one-point compactification) of an arbitrary topological space `X` and prove some properties inherited from `X`. diff --git a/src/topology/algebra/mul_action.lean b/src/topology/algebra/mul_action.lean index e0108ac9622d1..95c3c5e9da58a 100644 --- a/src/topology/algebra/mul_action.lean +++ b/src/topology/algebra/mul_action.lean @@ -11,6 +11,9 @@ import topology.algebra.const_mul_action /-! # Continuous monoid action +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define class `has_continuous_smul`. We say `has_continuous_smul M X` if `M` acts on `X` and the map `(c, x) ↦ c • x` is continuous on `M × X`. We reuse this class for topological (semi)modules, vector spaces and algebras. diff --git a/src/topology/algebra/order/left_right_lim.lean b/src/topology/algebra/order/left_right_lim.lean index 6628268721fd9..17a1b5487f5af 100644 --- a/src/topology/algebra/order/left_right_lim.lean +++ b/src/topology/algebra/order/left_right_lim.lean @@ -9,6 +9,9 @@ import topology.algebra.order.left_right /-! # Left and right limits +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define the (strict) left and right limits of a function. * `left_lim f x` is the strict left limit of `f` at `x` (using `f x` as a garbage value if `x` diff --git a/src/topology/quasi_separated.lean b/src/topology/quasi_separated.lean index b4ad41e647c87..bbd56a645a78d 100644 --- a/src/topology/quasi_separated.lean +++ b/src/topology/quasi_separated.lean @@ -10,6 +10,9 @@ import topology.noetherian_space /-! # Quasi-separated spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A topological space is quasi-separated if the intersections of any pairs of compact open subsets are still compact. Notable examples include spectral spaces, Noetherian spaces, and Hausdorff spaces. diff --git a/src/topology/sets/order.lean b/src/topology/sets/order.lean index d44b1dc3f74e4..b11ca9821a8d3 100644 --- a/src/topology/sets/order.lean +++ b/src/topology/sets/order.lean @@ -9,6 +9,9 @@ import topology.sets.closeds /-! # Clopen upper sets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define the type of clopen upper sets. -/ From a231f964b9ec8d1e12a28326b556123258a52829 Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Tue, 14 Feb 2023 15:23:31 +0000 Subject: [PATCH 0476/1291] feat(analysis/fourier): Basic framework for Fourier transforms (#18428) This adds a file whose purpose is to define the Fourier transform in very great generality -- this is a spinoff from #18392. --- src/analysis/fourier/fourier_transform.lean | 264 ++++++++++++++++++ .../fourier/riemann_lebesgue_lemma.lean | 42 ++- 2 files changed, 304 insertions(+), 2 deletions(-) create mode 100644 src/analysis/fourier/fourier_transform.lean diff --git a/src/analysis/fourier/fourier_transform.lean b/src/analysis/fourier/fourier_transform.lean new file mode 100644 index 0000000000000..2393022800122 --- /dev/null +++ b/src/analysis/fourier/fourier_transform.lean @@ -0,0 +1,264 @@ +/- +Copyright (c) 2023 David Loeffler. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: David Loeffler +-/ + +import analysis.special_functions.complex.circle +import measure_theory.group.integration +import measure_theory.integral.integral_eq_improper + +/-! +# The Fourier transform + +We set up the Fourier transform for complex-valued functions on finite-dimensional spaces. + +## Design choices + +In namespace `vector_fourier`, we define the Fourier integral in the following context: +* `𝕜` is a commutative ring. +* `V` and `W` are `𝕜`-modules. +* `e` is a unitary additive character of `𝕜`, i.e. a homomorphism `(multiplicative 𝕜) →* circle`. +* `μ` is a measure on `V`. +* `L` is a `𝕜`-bilinear form `V × W → 𝕜`. +* `E` is a complete normed `ℂ`-vector space. + +With these definitions, we define `fourier_integral` to be the map from functions `V → E` to +functions `W → E` that sends `f` to + +`λ w, ∫ v in V, e [-L v w] • f v ∂μ`, + +where `e [x]` is notational sugar for `(e (multiplicative.of_add x) : ℂ)` (available in locale +`fourier_transform`). This includes the cases `W` is the dual of `V` and `L` is the canonical +pairing, or `W = V` and `L` is a bilinear form (e.g. an inner product). + +In namespace `fourier`, we consider the more familiar special case when `V = W = 𝕜` and `L` is the +multiplication map (but still allowing `𝕜` to be an arbitrary ring equipped with a measure). + +The most familiar case of all is when `V = W = 𝕜 = ℝ`, `L` is multiplication, `μ` is volume, and +`e` is `real.fourier_char`, i.e. the character `λ x, exp ((2 * π * x) * I)`. The Fourier integral +in this case is defined as `real.fourier_integral`. + +## Main results + +At present the only nontrivial lemma we prove is `continuous_fourier_integral`, stating that the +Fourier transform of an integrable function is continuous (under mild assumptions). +-/ + +noncomputable theory + +local notation `𝕊` := circle + +open measure_theory filter + +open_locale topology + +-- To avoid messing around with multiplicative vs. additive characters, we make a notation. +localized "notation e `[` x `]` := (e (multiplicative.of_add x) : ℂ)" in fourier_transform + +/-! ## Fourier theory for functions on general vector spaces -/ +namespace vector_fourier + +variables + {𝕜 : Type*} [comm_ring 𝕜] + {V : Type*} [add_comm_group V] [module 𝕜 V] [measurable_space V] + {W : Type*} [add_comm_group W] [module 𝕜 W] + {E : Type*} [normed_add_comm_group E] [normed_space ℂ E] + +section defs + +variables [complete_space E] + +/-- The Fourier transform integral for `f : V → E`, with respect to a bilinear form `L : V × W → 𝕜` +and an additive character `e`. -/ +def fourier_integral + (e : (multiplicative 𝕜) →* 𝕊) (μ : measure V) (L : V →ₗ[𝕜] W →ₗ[𝕜] 𝕜) + (f : V → E) (w : W) : E := +∫ v, e [-L v w] • f v ∂μ + +lemma fourier_integral_smul_const + (e : (multiplicative 𝕜) →* 𝕊) (μ : measure V) (L : V →ₗ[𝕜] W →ₗ[𝕜] 𝕜) + (f : V → E) (r : ℂ) : + fourier_integral e μ L (r • f) = r • (fourier_integral e μ L f) := +begin + ext1 w, + simp only [pi.smul_apply, fourier_integral, smul_comm _ r, integral_smul], +end + +/-- The uniform norm of the Fourier integral of `f` is bounded by the `L¹` norm of `f`. -/ +lemma fourier_integral_norm_le (e : (multiplicative 𝕜) →* 𝕊) {μ : measure V} + (L : V →ₗ[𝕜] W →ₗ[𝕜] 𝕜) {f : V → E} (hf : integrable f μ) (w : W) : + ‖fourier_integral e μ L f w‖ ≤ ‖hf.to_L1 f‖ := +begin + rw L1.norm_of_fun_eq_integral_norm, + refine (norm_integral_le_integral_norm _).trans (le_of_eq _), + simp_rw [norm_smul, complex.norm_eq_abs, abs_coe_circle, one_mul], +end + +/-- The Fourier integral converts right-translation into scalar multiplication by a phase factor.-/ +lemma fourier_integral_comp_add_right [has_measurable_add V] + (e : (multiplicative 𝕜) →* 𝕊) (μ : measure V) [μ.is_add_right_invariant] + (L : V →ₗ[𝕜] W →ₗ[𝕜] 𝕜) (f : V → E) (v₀ : V) : + fourier_integral e μ L (f ∘ (λ v, v + v₀)) = λ w, e [L v₀ w] • fourier_integral e μ L f w := +begin + ext1 w, + dsimp only [fourier_integral, function.comp_apply], + conv in (L _) { rw ←add_sub_cancel v v₀ }, + rw integral_add_right_eq_self (λ (v : V), e [-L (v - v₀) w] • f v), + swap, apply_instance, + dsimp only, + rw ←integral_smul, + congr' 1 with v, + rw [←smul_assoc, smul_eq_mul, ←submonoid.coe_mul, ←e.map_mul, ←of_add_add, ←linear_map.neg_apply, + ←sub_eq_add_neg, ←linear_map.sub_apply, linear_map.map_sub, neg_sub] +end + +end defs + +section continuous +/- In this section we assume 𝕜, V, W have topologies, and L, e are continuous (but f needn't be). + This is used to ensure that `e [-L v w]` is (ae strongly) measurable. We could get away with + imposing only a measurable-space structure on 𝕜 (it doesn't have to be the Borel sigma-algebra of + a topology); but it seems hard to imagine cases where this extra generality would be useful, and + allowing it would complicate matters in the most important use cases. +-/ + +variables [topological_space 𝕜] [topological_ring 𝕜] [topological_space V] [borel_space V] + [topological_space W] {e : (multiplicative 𝕜) →* 𝕊} {μ : measure V} {L : V →ₗ[𝕜] W →ₗ[𝕜] 𝕜} + +/-- If `f` is integrable, then the Fourier integral is convergent for all `w`. -/ +lemma fourier_integral_convergent + (he : continuous e) (hL : continuous (λ p : V × W, L p.1 p.2)) + {f : V → E} (hf : integrable f μ) (w : W) : + integrable (λ (v : V), (e [-L v w]) • f v) μ := +begin + rw continuous_induced_rng at he, + have c : continuous (λ v, e[-L v w]), + { refine he.comp (continuous_of_add.comp (continuous.neg _)), + exact hL.comp (continuous_prod_mk.mpr ⟨continuous_id, continuous_const⟩) }, + rw ←integrable_norm_iff (c.ae_strongly_measurable.smul hf.1), + convert hf.norm, + ext1 v, + rw [norm_smul, complex.norm_eq_abs, abs_coe_circle, one_mul] +end + +variables [complete_space E] + +lemma fourier_integral_add + (he : continuous e) (hL : continuous (λ p : V × W, L p.1 p.2)) + {f g : V → E} (hf : integrable f μ) (hg : integrable g μ) : + (fourier_integral e μ L f) + (fourier_integral e μ L g) = fourier_integral e μ L (f + g) := +begin + ext1 w, + dsimp only [pi.add_apply, fourier_integral], + simp_rw smul_add, + rw integral_add, + { exact fourier_integral_convergent he hL hf w }, + { exact fourier_integral_convergent he hL hg w }, +end + +/-- The Fourier integral of an `L^1` function is a continuous function. -/ +lemma fourier_integral_continuous [topological_space.first_countable_topology W] + (he : continuous e) (hL : continuous (λ p : V × W, L p.1 p.2)) + {f : V → E} (hf : integrable f μ) : + continuous (fourier_integral e μ L f) := +begin + apply continuous_of_dominated, + { exact λ w, (fourier_integral_convergent he hL hf w).1 }, + { refine λ w, ae_of_all _ (λ v, _), + { exact λ v, ‖f v‖ }, + { rw [norm_smul, complex.norm_eq_abs, abs_coe_circle, one_mul] } }, + { exact hf.norm }, + { rw continuous_induced_rng at he, + refine ae_of_all _ (λ v, (he.comp (continuous_of_add.comp _)).smul continuous_const), + refine (hL.comp (continuous_prod_mk.mpr ⟨continuous_const, continuous_id⟩)).neg } +end + +end continuous + +end vector_fourier + +/-! ## Fourier theory for functions on `𝕜` -/ +namespace fourier + +variables {𝕜 : Type*} [comm_ring 𝕜] [measurable_space 𝕜] + {E : Type*} [normed_add_comm_group E] [normed_space ℂ E] + +section defs + +variables [complete_space E] + +/-- The Fourier transform integral for `f : 𝕜 → E`, with respect to the measure `μ` and additive +character `e`. -/ +def fourier_integral + (e : (multiplicative 𝕜) →* 𝕊) (μ : measure 𝕜) (f : 𝕜 → E) (w : 𝕜) : E := +vector_fourier.fourier_integral e μ (linear_map.mul 𝕜 𝕜) f w + +lemma fourier_integral_def (e : (multiplicative 𝕜) →* 𝕊) (μ : measure 𝕜) (f : 𝕜 → E) (w : 𝕜) : + fourier_integral e μ f w = ∫ (v : 𝕜), e[-(v * w)] • f v ∂μ := +rfl + +lemma fourier_integral_smul_const + (e : (multiplicative 𝕜) →* 𝕊) (μ : measure 𝕜) (f : 𝕜 → E) (r : ℂ) : + fourier_integral e μ (r • f) = r • (fourier_integral e μ f) := +vector_fourier.fourier_integral_smul_const _ _ _ _ _ + +/-- The uniform norm of the Fourier transform of `f` is bounded by the `L¹` norm of `f`. -/ +lemma fourier_integral_norm_le (e : (multiplicative 𝕜) →* 𝕊) {μ : measure 𝕜} + {f : 𝕜 → E} (hf : integrable f μ) (w : 𝕜) : + ‖fourier_integral e μ f w‖ ≤ ‖hf.to_L1 f‖ := +vector_fourier.fourier_integral_norm_le _ _ _ _ + +/-- The Fourier transform converts right-translation into scalar multiplication by a phase factor.-/ +lemma fourier_integral_comp_add_right [has_measurable_add 𝕜] + (e : (multiplicative 𝕜) →* 𝕊) (μ : measure 𝕜) [μ.is_add_right_invariant] (f : 𝕜 → E) (v₀ : 𝕜) : + fourier_integral e μ (f ∘ (λ v, v + v₀)) = λ w, e [v₀ * w] • fourier_integral e μ f w := +vector_fourier.fourier_integral_comp_add_right _ _ _ _ _ + +end defs + +end fourier + +open_locale real + +namespace real + +/-- The standard additive character of `ℝ`, given by `λ x, exp (2 * π * x * I)`. -/ +def fourier_char : (multiplicative ℝ) →* 𝕊 := +{ to_fun := λ z, exp_map_circle (2 * π * z.to_add), + map_one' := by rw [to_add_one, mul_zero, exp_map_circle_zero], + map_mul' := λ x y, by rw [to_add_mul, mul_add, exp_map_circle_add] } + +lemma fourier_char_apply (x : ℝ) : + real.fourier_char [x] = complex.exp (↑(2 * π * x) * complex.I) := +by refl + +@[continuity] +lemma continuous_fourier_char : continuous real.fourier_char := +(map_continuous exp_map_circle).comp (continuous_const.mul continuous_to_add) + +variables {E : Type*} [normed_add_comm_group E] [complete_space E] [normed_space ℂ E] + +lemma vector_fourier_integral_eq_integral_exp_smul + {V : Type*} [add_comm_group V] [module ℝ V] [measurable_space V] + {W : Type*} [add_comm_group W] [module ℝ W] + (L : V →ₗ[ℝ] W →ₗ[ℝ] ℝ) (μ : measure V) (f : V → E) (w : W) : + vector_fourier.fourier_integral fourier_char μ L f w + = ∫ (v : V), complex.exp (↑(-2 * π * L v w) * complex.I) • f v ∂μ := +by simp_rw [vector_fourier.fourier_integral, real.fourier_char_apply, mul_neg, neg_mul] + +/-- The Fourier integral for `f : ℝ → E`, with respect to the standard additive character and +measure on `ℝ`. -/ +def fourier_integral (f : ℝ → E) (w : ℝ) := fourier.fourier_integral fourier_char volume f w + +lemma fourier_integral_def (f : ℝ → E) (w : ℝ) : + fourier_integral f w = ∫ (v : ℝ), fourier_char [-(v * w)] • f v := +rfl + +lemma fourier_integral_eq_integral_exp_smul + {E : Type*} [normed_add_comm_group E] [complete_space E] [normed_space ℂ E] + (f : ℝ → E) (w : ℝ) : + fourier_integral f w = ∫ (v : ℝ), complex.exp (↑(-2 * π * v * w) * complex.I) • f v := +by simp_rw [fourier_integral_def, real.fourier_char_apply, mul_neg, neg_mul, mul_assoc] + +end real diff --git a/src/analysis/fourier/riemann_lebesgue_lemma.lean b/src/analysis/fourier/riemann_lebesgue_lemma.lean index 4046ebd59f3f0..2b2d53dd53313 100644 --- a/src/analysis/fourier/riemann_lebesgue_lemma.lean +++ b/src/analysis/fourier/riemann_lebesgue_lemma.lean @@ -3,12 +3,12 @@ Copyright (c) 2022 David Loeffler. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: David Loeffler -/ -import measure_theory.function.lp_space + import measure_theory.function.continuous_map_dense -import measure_theory.integral.interval_integral import measure_theory.integral.integral_eq_improper import measure_theory.group.integration import topology.continuous_function.zero_at_infty +import analysis.fourier.fourier_transform /-! # The Riemann-Lebesgue Lemma @@ -142,4 +142,42 @@ begin exact lt_add_of_pos_left _ hε, end +lemma tendsto_integral_mul_exp_at_bot_of_continuous_compact_support + (hf1 : continuous f) (hf2 : has_compact_support f) : + tendsto (λ t:ℝ, ∫ x:ℝ, exp (↑(t * x) * I) • f x) at_bot (𝓝 0) := +begin + have hg2 : has_compact_support (f ∘ has_neg.neg), + by simpa only [neg_one_smul] using hf2.comp_smul (neg_ne_zero.mpr $ one_ne_zero' ℝ), + convert (tendsto_integral_mul_exp_at_top_of_continuous_compact_support (hf1.comp continuous_neg) + hg2).comp tendsto_neg_at_bot_at_top, + ext1 t, + simp_rw [function.comp_app, neg_mul, ←mul_neg], + rw ←integral_neg_eq_self, +end + +lemma zero_at_infty_integral_mul_exp_of_continuous_compact_support + (hf1 : continuous f) (hf2 : has_compact_support f) : + tendsto (λ t:ℝ, ∫ x:ℝ, exp (↑(t * x) * I) • f x) (cocompact ℝ) (𝓝 0) := +begin + rw [real.cocompact_eq, tendsto_sup], + exact ⟨tendsto_integral_mul_exp_at_bot_of_continuous_compact_support hf1 hf2, + tendsto_integral_mul_exp_at_top_of_continuous_compact_support hf1 hf2⟩ +end + +open_locale fourier_transform + +/-- Riemann-Lebesgue lemma for continuous compactly-supported functions: the Fourier transform +tends to 0 at infinity. -/ +lemma real.fourier_integral_zero_at_infty_of_continuous_compact_support + (hc : continuous f) (hs : has_compact_support f) : + tendsto (real.fourier_integral f) (cocompact ℝ) (𝓝 0) := +begin + refine ((zero_at_infty_integral_mul_exp_of_continuous_compact_support hc hs).comp + (tendsto_cocompact_mul_left₀ + (mul_ne_zero (neg_ne_zero.mpr two_ne_zero) real.pi_pos.ne'))).congr (λ w, _), + rw [real.fourier_integral_eq_integral_exp_smul, function.comp_app], + congr' 1 with x:1, + ring_nf, +end + end continuous_compact_support From 369525b73f229ccd76a6ec0e0e0bf2be57599768 Mon Sep 17 00:00:00 2001 From: Yakov Pechersky Date: Tue, 14 Feb 2023 16:30:16 +0000 Subject: [PATCH 0477/1291] chore(algebra/group/units): is_add_unit.decidable (#18439) Missing `to_additive` I had actually added this to mathlib4 3 months ago in leanprover-community/mathlib4@cbc083419c0aeee4e99154f304f5f8ac66bce0d8 (leanprover-community/mathlib4#549), and never noticed it wasn't in mathlib3. mathlib4 hash-update placeholder pr: https://github.com/leanprover-community/mathlib4/pull/2278 --- src/algebra/group/units.lean | 1 + 1 file changed, 1 insertion(+) diff --git a/src/algebra/group/units.lean b/src/algebra/group/units.lean index 67e2461a08917..d22ea64c3f1d4 100644 --- a/src/algebra/group/units.lean +++ b/src/algebra/group/units.lean @@ -458,6 +458,7 @@ lemma coe_inv_mul (h : is_unit a) : ↑(h.unit)⁻¹ * a = 1 := units.mul_inv _ by convert h.unit.mul_inv /-- `is_unit x` is decidable if we can decide if `x` comes from `Mˣ`. -/ +@[to_additive "`is_add_unit x` is decidable if we can decide if `x` comes from `add_units M"] instance (x : M) [h : decidable (∃ u : Mˣ, ↑u = x)] : decidable (is_unit x) := h @[to_additive] lemma mul_left_inj (h : is_unit a) : b * a = c * a ↔ b = c := From 1ead22342e1a078bd44744ace999f85756555d35 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Wed, 15 Feb 2023 04:29:49 +0000 Subject: [PATCH 0478/1291] chore(*): add mathlib4 synchronization comments (#18443) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.module.submodule.pointwise` * `category_theory.comma` * `category_theory.products.associator` * `category_theory.products.basic` * `category_theory.products.bifunctor` * `combinatorics.partition` * `data.finsupp.lex` * `data.nat.count` * `data.nat.periodic` * `group_theory.index` * `linear_algebra.pi` * `linear_algebra.span` * `number_theory.ADE_inequality` * `topology.algebra.monoid` * `topology.sets.compacts` * `topology.uniform_space.compact_convergence` --- src/algebra/module/submodule/pointwise.lean | 3 +++ src/category_theory/comma.lean | 3 +++ src/category_theory/products/associator.lean | 3 +++ src/category_theory/products/basic.lean | 3 +++ src/category_theory/products/bifunctor.lean | 3 +++ src/combinatorics/partition.lean | 3 +++ src/data/finsupp/lex.lean | 3 +++ src/data/nat/count.lean | 3 +++ src/data/nat/periodic.lean | 3 +++ src/group_theory/index.lean | 3 +++ src/linear_algebra/pi.lean | 3 +++ src/linear_algebra/span.lean | 3 +++ src/number_theory/ADE_inequality.lean | 3 +++ src/topology/algebra/monoid.lean | 3 +++ src/topology/sets/compacts.lean | 3 +++ src/topology/uniform_space/compact_convergence.lean | 3 +++ 16 files changed, 48 insertions(+) diff --git a/src/algebra/module/submodule/pointwise.lean b/src/algebra/module/submodule/pointwise.lean index b119c8f629ff1..bc4be0638ac43 100644 --- a/src/algebra/module/submodule/pointwise.lean +++ b/src/algebra/module/submodule/pointwise.lean @@ -8,6 +8,9 @@ import linear_algebra.span /-! # Pointwise instances on `submodule`s +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file provides: * `submodule.has_pointwise_neg` diff --git a/src/category_theory/comma.lean b/src/category_theory/comma.lean index 15215ad7ea483..dd632fe0d48bd 100644 --- a/src/category_theory/comma.lean +++ b/src/category_theory/comma.lean @@ -10,6 +10,9 @@ import category_theory.eq_to_hom /-! # Comma categories +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A comma category is a construction in category theory, which builds a category out of two functors with a common codomain. Specifically, for functors `L : A ⥤ T` and `R : B ⥤ T`, an object in `comma L R` is a morphism `hom : L.obj left ⟶ R.obj right` for some objects `left : A` and diff --git a/src/category_theory/products/associator.lean b/src/category_theory/products/associator.lean index 429dbc51729a2..f1b340d0b8750 100644 --- a/src/category_theory/products/associator.lean +++ b/src/category_theory/products/associator.lean @@ -6,6 +6,9 @@ Authors: Stephen Morgan, Scott Morrison import category_theory.products.basic /-! +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The associator functor `((C × D) × E) ⥤ (C × (D × E))` and its inverse form an equivalence. -/ diff --git a/src/category_theory/products/basic.lean b/src/category_theory/products/basic.lean index 65d15df5afe92..8fcb51d0ec542 100644 --- a/src/category_theory/products/basic.lean +++ b/src/category_theory/products/basic.lean @@ -10,6 +10,9 @@ import data.prod.basic /-! # Cartesian products of categories +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define the category instance on `C × D` when `C` and `D` are categories. We define: diff --git a/src/category_theory/products/bifunctor.lean b/src/category_theory/products/bifunctor.lean index f8b62a738ec1f..f20ec369fff1a 100644 --- a/src/category_theory/products/bifunctor.lean +++ b/src/category_theory/products/bifunctor.lean @@ -7,6 +7,9 @@ import category_theory.products.basic /-! # Lemmas about functors out of product categories. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ open category_theory diff --git a/src/combinatorics/partition.lean b/src/combinatorics/partition.lean index da174e729cfae..40f4e5e868093 100644 --- a/src/combinatorics/partition.lean +++ b/src/combinatorics/partition.lean @@ -11,6 +11,9 @@ import tactic.apply_fun /-! # Partitions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A partition of a natural number `n` is a way of writing `n` as a sum of positive integers, where the order does not matter: two sums that differ only in the order of their summands are considered the same partition. This notion is closely related to that of a composition of `n`, but in a composition diff --git a/src/data/finsupp/lex.lean b/src/data/finsupp/lex.lean index 48933ca201879..09fdd2c4fc9ef 100644 --- a/src/data/finsupp/lex.lean +++ b/src/data/finsupp/lex.lean @@ -10,6 +10,9 @@ import data.finsupp.to_dfinsupp /-! # Lexicographic order on finitely supported functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the lexicographic order on `finsupp`. -/ diff --git a/src/data/nat/count.lean b/src/data/nat/count.lean index a35fd9efa9ee8..3ec1de6d907b2 100644 --- a/src/data/nat/count.lean +++ b/src/data/nat/count.lean @@ -9,6 +9,9 @@ import tactic.ring /-! # Counting on ℕ +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the `count` function, which gives, for any predicate on the natural numbers, "how many numbers under `k` satisfy this predicate?". We then prove several expected lemmas about `count`, relating it to the cardinality of other diff --git a/src/data/nat/periodic.lean b/src/data/nat/periodic.lean index 501e32e15a159..48eba2c749525 100644 --- a/src/data/nat/periodic.lean +++ b/src/data/nat/periodic.lean @@ -10,6 +10,9 @@ import data.nat.interval /-! # Periodic Functions on ℕ +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file identifies a few functions on `ℕ` which are periodic, and also proves a lemma about periodic predicates which helps determine their cardinality when filtering intervals over them. -/ diff --git a/src/group_theory/index.lean b/src/group_theory/index.lean index bb1ee6f8e7064..ce2d7fe11c151 100644 --- a/src/group_theory/index.lean +++ b/src/group_theory/index.lean @@ -11,6 +11,9 @@ import group_theory.group_action.quotient /-! # Index of a Subgroup +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define the index of a subgroup, and prove several divisibility properties. Several theorems proved in this file are known as Lagrange's theorem. diff --git a/src/linear_algebra/pi.lean b/src/linear_algebra/pi.lean index bec285f856519..d7126156e85c6 100644 --- a/src/linear_algebra/pi.lean +++ b/src/linear_algebra/pi.lean @@ -9,6 +9,9 @@ import logic.equiv.fin /-! # Pi types of modules +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines constructors for linear maps whose domains or codomains are pi types. It contains theorems relating these to each other, as well as to `linear_map.ker`. diff --git a/src/linear_algebra/span.lean b/src/linear_algebra/span.lean index 004f7e0b618bf..54c850ab6fabd 100644 --- a/src/linear_algebra/span.lean +++ b/src/linear_algebra/span.lean @@ -11,6 +11,9 @@ import order.omega_complete_partial_order /-! # The span of a set of vectors, as a submodule +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + * `submodule.span s` is defined to be the smallest submodule containing the set `s`. ## Notations diff --git a/src/number_theory/ADE_inequality.lean b/src/number_theory/ADE_inequality.lean index 67760b10cddb9..5ba14070ab98c 100644 --- a/src/number_theory/ADE_inequality.lean +++ b/src/number_theory/ADE_inequality.lean @@ -16,6 +16,9 @@ import tactic.interval_cases /-! # The inequality `p⁻¹ + q⁻¹ + r⁻¹ > 1` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we classify solutions to the inequality `(p⁻¹ + q⁻¹ + r⁻¹ : ℚ) > 1`, for positive natural numbers `p`, `q`, and `r`. diff --git a/src/topology/algebra/monoid.lean b/src/topology/algebra/monoid.lean index c430df3cdbc22..fa433e34f9574 100644 --- a/src/topology/algebra/monoid.lean +++ b/src/topology/algebra/monoid.lean @@ -11,6 +11,9 @@ import algebra.big_operators.pi /-! # Theory of topological monoids +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define mixin classes `has_continuous_mul` and `has_continuous_add`. While in many applications the underlying type is a monoid (multiplicative or additive), we do not require this in the definitions. diff --git a/src/topology/sets/compacts.lean b/src/topology/sets/compacts.lean index 28b0a6f2f6df4..e4a4b96b9fb58 100644 --- a/src/topology/sets/compacts.lean +++ b/src/topology/sets/compacts.lean @@ -9,6 +9,9 @@ import topology.quasi_separated /-! # Compact sets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define a few types of compact sets in a topological space. ## Main Definitions diff --git a/src/topology/uniform_space/compact_convergence.lean b/src/topology/uniform_space/compact_convergence.lean index 044ef05e4be37..f10663527d9b2 100644 --- a/src/topology/uniform_space/compact_convergence.lean +++ b/src/topology/uniform_space/compact_convergence.lean @@ -9,6 +9,9 @@ import topology.uniform_space.uniform_convergence /-! # Compact convergence (uniform convergence on compact sets) +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given a topological space `α` and a uniform space `β` (e.g., a metric space or a topological group), the space of continuous maps `C(α, β)` carries a natural uniform space structure. We define this uniform space structure in this file and also prove the following properties of the topology it From 372edc36e5d2caafdd135769e0136b5a59186834 Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Wed, 15 Feb 2023 07:25:40 +0000 Subject: [PATCH 0479/1291] feat (analysis/special_functions): limit formula and reflection formula for Gamma (#18404) This adds Euler's limit formula and the reflection formula for the Gamma function. As a consequence we get that Gamma does not vanish on the complex plane (except at non-positive integers, where it is undefined, and we have arbitrarily chosen it to be 0). --- src/analysis/normed/field/basic.lean | 5 + src/analysis/special_functions/gamma.lean | 415 +++++++++++++++++++--- src/analysis/specific_limits/basic.lean | 29 ++ src/data/complex/exponential.lean | 12 + 4 files changed, 413 insertions(+), 48 deletions(-) diff --git a/src/analysis/normed/field/basic.lean b/src/analysis/normed/field/basic.lean index efa11872c9429..5a559308339d4 100644 --- a/src/analysis/normed/field/basic.lean +++ b/src/analysis/normed/field/basic.lean @@ -487,6 +487,11 @@ begin simp, end +/-- A normed division ring is a topological division ring. -/ +@[priority 100] -- see Note [lower instance priority] +instance normed_division_ring.to_topological_division_ring : topological_division_ring α := +{ } + lemma norm_map_one_of_pow_eq_one [monoid β] (φ : β →* α) {x : β} {k : ℕ+} (h : x ^ (k : ℕ) = 1) : ‖φ x‖ = 1 := diff --git a/src/analysis/special_functions/gamma.lean b/src/analysis/special_functions/gamma.lean index 3bc21785cf2eb..358b1b428b146 100644 --- a/src/analysis/special_functions/gamma.lean +++ b/src/analysis/special_functions/gamma.lean @@ -7,6 +7,7 @@ import measure_theory.integral.exp_decay import analysis.calculus.parametric_integral import analysis.special_functions.integrals import analysis.convolution +import analysis.special_functions.trigonometric.euler_sine_prod /-! # The Gamma and Beta functions @@ -17,35 +18,35 @@ integral `Γ(s) = ∫ x in Ioi 0, exp (-x) * x ^ (s - 1)` in the range where thi We show that this integral satisfies `Γ(1) = 1` and `Γ(s + 1) = s * Γ(s)`; hence we can define `Γ(s)` for all `s` as the unique function satisfying this recurrence and agreeing with Euler's -integral in the convergence range. In the complex case we also prove that the resulting function is -holomorphic on `ℂ` away from the points `{-n : n ∈ ℕ}`. +integral in the convergence range. (If `s = -n` for `n ∈ ℕ`, then the function is undefined, and we +set it to be `0` by convention.) -## Gamma function: main statements (real case) +## Gamma function: main statements (complex case) -* `real.Gamma` : the `Γ` function (of a real variable). -* `real.Gamma_eq_integral` : for `0 < s`, `Γ(s)` agrees with Euler's integral - `∫ (x:ℝ) in Ioi 0, exp (-x) * x ^ (s - 1)` -* `real.Gamma_add_one` : for all `s : ℝ` with `s ≠ 0`, we have `Γ(s + 1) = s Γ(s)`. -* `real.Gamma_nat_eq_factorial` : for all `n : ℕ` we have `Γ (n + 1) = n!`. -* `real.differentiable_at_Gamma` : `Γ` is real-differentiable at all `s : ℝ` with +* `complex.Gamma`: the `Γ` function (of a complex variable). +* `complex.Gamma_eq_integral`: for `0 < re s`, `Γ(s)` agrees with Euler's integral. +* `complex.Gamma_add_one`: for all `s : ℂ` with `s ≠ 0`, we have `Γ (s + 1) = s Γ(s)`. +* `complex.Gamma_nat_eq_factorial`: for all `n : ℕ` we have `Γ (n + 1) = n!`. +* `complex.differentiable_at_Gamma`: `Γ` is complex-differentiable at all `s : ℂ` with `s ∉ {-n : n ∈ ℕ}`. -* `real.Gamma_ne_zero`: for all `s : ℝ` with `s ∉ {-n : n ∈ ℕ}` we have `Γ s ≠ 0`. -* `real.tendsto_log_Gamma`: for all `0 < s`, the limit as `n → ∞` of the sequence - `n ↦ s log n + log n! - (log s + ... + log (s + n))` is `log Γ(s)`. +* `complex.Gamma_ne_zero`: for all `s : ℂ` with `s ∉ {-n : n ∈ ℕ}` we have `Γ s ≠ 0`. +* `complex.Gamma_seq_tendsto_Gamma`: for all `s`, the limit as `n → ∞` of the sequence + `n ↦ n ^ s * n! / (s * (s + 1) * ... * (s + n))` is `Γ(s)`. +* `complex.Gamma_mul_Gamma_one_sub`: Euler's reflection formula + `Gamma s * Gamma (1 - s) = π / sin π s`. + +## Gamma function: main statements (real case) + +* `real.Gamma`: the `Γ` function (of a real variable). +* Real counterparts of all the properties of the complex Gamma function listed above: + `real.Gamma_eq_integral`, `real.Gamma_add_one`, `real.Gamma_nat_eq_factorial`, + `real.differentiable_at_Gamma`, `real.Gamma_ne_zero`, `real.Gamma_seq_tendsto_Gamma`, + `real.Gamma_mul_Gamma_one_sub`. * `real.convex_on_log_Gamma` : `log ∘ Γ` is convex on `Ioi 0`. * `real.eq_Gamma_of_log_convex` : the Bohr-Mollerup theorem, which states that the `Γ` function is the unique log-convex, positive-valued function on `Ioi 0` satisfying the functional equation and having `Γ 1 = 1`. -## Gamma function: main statements (complex case) - -* `complex.Gamma` : the `Γ` function (of a complex variable). -* `complex.Gamma_eq_integral` : for `0 < re s`, `Γ(s)` agrees with Euler's integral. -* `complex.Gamma_add_one` : for all `s : ℂ` with `s ≠ 0`, we have `Γ(s + 1) = s Γ(s)`. -* `complex.Gamma_nat_eq_factorial` : for all `n : ℕ` we have `Γ (n + 1) = n!`. -* `complex.differentiable_at_Gamma` : `Γ` is complex-differentiable at all `s : ℂ` with - `s ∉ {-n : n ∈ ℕ}`. - ## Beta function * `complex.beta_integral`: the Beta function `Β(u, v)`, where `u`, `v` are complex with positive @@ -373,6 +374,24 @@ begin simp only [nat.cast_succ, nat.factorial_succ, nat.cast_mul], congr, exact hn }, end +/-- At `0` the Gamma function is undefined; by convention we assign it the value `0`. -/ +lemma Gamma_zero : Gamma 0 = 0 := +by simp_rw [Gamma, zero_re, sub_zero, nat.floor_one, Gamma_aux, div_zero] + +/-- At `-n` for `n ∈ ℕ`, the Gamma function is undefined; by convention we assign it the value 0. -/ +lemma Gamma_neg_nat_eq_zero (n : ℕ) : Gamma (-n) = 0 := +begin + induction n with n IH, + { rw [nat.cast_zero, neg_zero, Gamma_zero] }, + { have A : -(n.succ : ℂ) ≠ 0, + { rw [neg_ne_zero, nat.cast_ne_zero], + apply nat.succ_ne_zero }, + have : -(n:ℂ) = -↑n.succ + 1, by simp, + rw [this, Gamma_add_one _ A] at IH, + contrapose! IH, + exact mul_ne_zero A IH } +end + lemma Gamma_conj (s : ℂ) : Gamma (conj s) = conj (Gamma s) := begin suffices : ∀ (n:ℕ) (s:ℂ) , Gamma_aux n (conj s) = conj (Gamma_aux n s), from this _ _, @@ -530,7 +549,7 @@ begin (Gamma_integral_convergent (zero_lt_one.trans hs)) hF'_meas h_bound bound_integrable h_diff), end -lemma differentiable_at_Gamma_aux (s : ℂ) (n : ℕ) (h1 : (1 - s.re) < n ) (h2 : ∀ m:ℕ, s + m ≠ 0) : +lemma differentiable_at_Gamma_aux (s : ℂ) (n : ℕ) (h1 : (1 - s.re) < n ) (h2 : ∀ m : ℕ, s ≠ -m) : differentiable_at ℂ (Gamma_aux n) s := begin induction n with n hn generalizing s, @@ -540,13 +559,16 @@ begin specialize hn (s + 1), have a : 1 - (s + 1).re < ↑n, { rw nat.cast_succ at h1, rw [complex.add_re, complex.one_re], linarith }, - have b : ∀ m:ℕ, s + 1 + m ≠ 0, - { intro m, have := h2 (1 + m), rwa [nat.cast_add, nat.cast_one, ←add_assoc] at this }, + have b : ∀ m : ℕ, s + 1 ≠ -m, + { intro m, have := h2 (1 + m), + contrapose! this, + rw ←eq_sub_iff_add_eq at this, + simpa using this }, refine differentiable_at.div (differentiable_at.comp _ (hn a b) _) _ _, simp, simp, simpa using h2 0 } end -theorem differentiable_at_Gamma (s : ℂ) (hs : ∀ m:ℕ, s + m ≠ 0) : differentiable_at ℂ Gamma s := +theorem differentiable_at_Gamma (s : ℂ) (hs : ∀ m : ℕ, s ≠ -m) : differentiable_at ℂ Gamma s := begin let n := ⌊1 - s.re⌋₊ + 1, have hn : 1 - s.re < n := by exact_mod_cast nat.lt_floor_add_one (1 - s.re), @@ -604,6 +626,19 @@ theorem Gamma_nat_eq_factorial (n : ℕ) : Gamma (n + 1) = n! := by rw [Gamma, complex.of_real_add, complex.of_real_nat_cast, complex.of_real_one, complex.Gamma_nat_eq_factorial, ←complex.of_real_nat_cast, complex.of_real_re] +/-- At `0` the Gamma function is undefined; by convention we assign it the value `0`. -/ +lemma Gamma_zero : Gamma 0 = 0 := +by simpa only [←complex.of_real_zero, complex.Gamma_of_real, complex.of_real_inj] + using complex.Gamma_zero + +/-- At `-n` for `n ∈ ℕ`, the Gamma function is undefined; by convention we assign it the value `0`. +-/ +lemma Gamma_neg_nat_eq_zero (n : ℕ) : Gamma (-n) = 0 := +begin + simpa only [←complex.of_real_nat_cast, ←complex.of_real_neg, complex.Gamma_of_real, + complex.of_real_eq_zero] using complex.Gamma_neg_nat_eq_zero n, +end + lemma Gamma_pos_of_pos {s : ℝ} (hs : 0 < s) : 0 < Gamma s := begin rw Gamma_eq_integral hs, @@ -620,7 +655,9 @@ begin { exact Gamma_integral_convergent hs }, end -lemma Gamma_ne_zero {s : ℝ} (hs : ∀ m:ℕ, s + m ≠ 0) : Gamma s ≠ 0 := +/-- The Gamma function does not vanish on `ℝ` (except at non-positive integers, where the function +is mathematically undefined and we set it to `0` by convention). -/ +lemma Gamma_ne_zero {s : ℝ} (hs : ∀ m : ℕ, s ≠ -m) : Gamma s ≠ 0 := begin suffices : ∀ {n : ℕ}, (-(n:ℝ) < s) → Gamma s ≠ 0, { apply this, @@ -636,9 +673,12 @@ begin have : Gamma (s + 1) ≠ 0, { apply n_ih, { intro m, - convert hs (1 + m) using 1, + specialize hs (1 + m), + contrapose! hs, + rw ←eq_sub_iff_add_eq at hs, + rw hs, push_cast, - ring, }, + ring }, { rw [nat.succ_eq_add_one, nat.cast_add, nat.cast_one, neg_add] at hs', linarith } }, rw [Gamma_add_one, mul_ne_zero_iff] at this, @@ -646,13 +686,13 @@ begin { simpa using hs 0 } }, end -lemma differentiable_at_Gamma {s : ℝ} (hs : ∀ m:ℕ, s + m ≠ 0) : differentiable_at ℝ Gamma s := +lemma Gamma_eq_zero_iff (s : ℝ) : Gamma s = 0 ↔ ∃ m : ℕ, s = -m := +⟨by { contrapose!, exact Gamma_ne_zero }, by { rintro ⟨m, rfl⟩, exact Gamma_neg_nat_eq_zero m }⟩ + +lemma differentiable_at_Gamma {s : ℝ} (hs : ∀ m : ℕ, s ≠ -m) : differentiable_at ℝ Gamma s := begin - apply has_deriv_at.differentiable_at, - apply has_deriv_at.real_of_complex, - apply differentiable_at.has_deriv_at, - apply complex.differentiable_at_Gamma, - simp_rw [←complex.of_real_nat_cast, ←complex.of_real_add, complex.of_real_ne_zero], + refine ((complex.differentiable_at_Gamma _ _).has_deriv_at).real_of_complex.differentiable_at, + simp_rw [←complex.of_real_nat_cast, ←complex.of_real_neg, ne.def, complex.of_real_inj], exact hs, end @@ -743,36 +783,40 @@ begin { rw convex_iff_is_preconnected, refine is_preconnected_Ioi.image _ (λ x hx, continuous_at.continuous_within_at _), refine (differentiable_at_Gamma (λ m, _)).continuous_at.log (Gamma_pos_of_pos hx).ne', - exact (add_pos_of_pos_of_nonneg hx (nat.cast_nonneg m)).ne' }, + exact (neg_lt_iff_pos_add.mpr (add_pos_of_pos_of_nonneg hx (nat.cast_nonneg m))).ne' } end section bohr_mollerup -/-! ## The Euler limit formula and the Bohr-Mollerup theorem +/-! ## The Bohr-Mollerup theorem In this section we prove two interrelated statements about the `Γ` function on the positive reals: -* the Euler limit formula `real.tendsto_log_gamma_seq`, stating that the sequence - `x * log n + log n! - ∑ (m : ℕ) in finset.range (n + 1), log (x + m)` +* the Euler limit formula `real.bohr_mollerup.tendsto_log_gamma_seq`, stating that for positive + real `x` the sequence `x * log n + log n! - ∑ (m : ℕ) in finset.range (n + 1), log (x + m)` tends to `log Γ(x)` as `n → ∞`. * the Bohr-Mollerup theorem (`real.eq_Gamma_of_log_convex`) which states that `Γ` is the unique *log-convex*, positive-real-valued function on the positive reals satisfying `f (x + 1) = x f x` and `f 1 = 1`. To do this, we prove that any function satisfying the hypotheses of the Bohr--Mollerup theorem must -agree with the limit in the Gauss formula, so there is at most one such function. Then we show that -`Γ` satisfies these conditions. +agree with the limit in the Euler limit formula, so there is at most one such function. Then we +show that `Γ` satisfies these conditions. + +Since most of the auxiliary lemmas for the Bohr-Mollerup theorem are of no relevance outside the +context of this proof, we place them in a separate namespace `real.bohr_mollerup` to avoid clutter. +(This includes the logarithmic form of the Euler limit formula, since later we will prove a more +general form of the Euler limit formula valid for any real or complex `x`; see +`real.Gamma_seq_tendsto_Gamma` and `complex.Gamma_seq_tendsto_Gamma`.) -/ +namespace bohr_mollerup + /-- The function `n ↦ x log n + log n! - (log x + ... + log (x + n))`, which we will show tends to `log (Gamma x)` as `n → ∞`. -/ def log_gamma_seq (x : ℝ) (n : ℕ) : ℝ := x * log n + log n! - ∑ (m : ℕ) in finset.range (n + 1), log (x + m) -/-! The following are auxiliary lemmas for the Bohr-Mollerup theorem, which are -placed in a separate namespace `bohr_mollerup` to avoid clutter. -/ -namespace bohr_mollerup - variables {f : ℝ → ℝ} {x : ℝ} {n : ℕ} lemma f_nat_eq (hf_feq : ∀ {y:ℝ}, 0 < y → f (y + 1) = f y + log y) (hn : n ≠ 0) : @@ -943,8 +987,6 @@ begin ring } }, end -end bohr_mollerup - lemma tendsto_log_Gamma {x : ℝ} (hx : 0 < x) : tendsto (log_gamma_seq x) at_top (𝓝 $ log (Gamma x)) := begin @@ -955,6 +997,8 @@ begin rw [function.comp_app, Gamma_add_one hy.ne', log_mul hy.ne' (Gamma_pos_of_pos hy).ne', add_comm], end +end bohr_mollerup -- (namespace) + /-- The **Bohr-Mollerup theorem**: the Gamma function is the *unique* log-convex, positive-valued function on the positive reals which satisfies `f 1 = 1` and `f (x + 1) = x * f x` for all `x`. -/ lemma eq_Gamma_of_log_convex {f : ℝ → ℝ} @@ -969,13 +1013,13 @@ begin intros x hx, have e1 := bohr_mollerup.tendsto_log_gamma_seq hf_conv _ hx, { rw [function.comp_app log f 1, hf_one, log_one, sub_zero] at e1, - exact tendsto_nhds_unique e1 (tendsto_log_Gamma hx) }, + exact tendsto_nhds_unique e1 (bohr_mollerup.tendsto_log_Gamma hx) }, { intros y hy, rw [function.comp_app, hf_feq hy, log_mul hy.ne' (hf_pos hy).ne'], ring } end -end bohr_mollerup +end bohr_mollerup -- (section) section strict_mono @@ -1018,6 +1062,8 @@ end real section beta_integral +/-! ## The Beta function -/ + namespace complex notation `cexp` := complex.exp @@ -1214,3 +1260,276 @@ end end complex end beta_integral + +section limit_formula + +/-! ## The Euler limit formula -/ + +namespace complex + +/-- The sequence with `n`-th term `n ^ s * n! / (s * (s + 1) * ... * (s + n))`, for complex `s`. +We will show that this tends to `Γ(s)` as `n → ∞`. -/ +noncomputable def Gamma_seq (s : ℂ) (n : ℕ) := +(n:ℂ) ^ s * n! / ∏ (j:ℕ) in finset.range (n + 1), (s + j) + +lemma Gamma_seq_eq_beta_integral_of_re_pos {s : ℂ} (hs : 0 < re s) (n : ℕ) : + Gamma_seq s n = n ^ s * beta_integral s (n + 1) := +by rw [Gamma_seq, beta_integral_eval_nat_add_one_right hs n, ←mul_div_assoc] + +lemma Gamma_seq_add_one_left (s : ℂ) {n : ℕ} (hn : n ≠ 0) : + (Gamma_seq (s + 1) n) / s = n / (n + 1 + s) * Gamma_seq s n := +begin + conv_lhs { rw [Gamma_seq, finset.prod_range_succ, div_div] }, + conv_rhs { rw [Gamma_seq, finset.prod_range_succ', nat.cast_zero, add_zero, div_mul_div_comm, + ←mul_assoc, ←mul_assoc, mul_comm _ (finset.prod _ _)] }, + congr' 3, + { rw [cpow_add _ _ (nat.cast_ne_zero.mpr hn), cpow_one, mul_comm] }, + { refine finset.prod_congr (by refl) (λ x hx, _), + push_cast, ring }, + { abel } +end + +lemma Gamma_seq_eq_approx_Gamma_integral {s : ℂ} (hs : 0 < re s) {n : ℕ} (hn : n ≠ 0) : + Gamma_seq s n = ∫ x:ℝ in 0..n, ↑((1 - x / n) ^ n) * (x:ℂ) ^ (s - 1) := +begin + have : ∀ (x : ℝ), x = x / n * n, by { intro x, rw div_mul_cancel, exact nat.cast_ne_zero.mpr hn }, + conv in (↑_ ^ _) { congr, rw this x }, + rw Gamma_seq_eq_beta_integral_of_re_pos hs, + rw [beta_integral, @interval_integral.integral_comp_div _ _ _ _ 0 n _ + (λ x, ↑((1 - x) ^ n) * ↑(x * ↑n) ^ (s - 1) : ℝ → ℂ) (nat.cast_ne_zero.mpr hn), + real_smul, zero_div, div_self, add_sub_cancel, ←interval_integral.integral_const_mul, + ←interval_integral.integral_const_mul], + swap, { exact nat.cast_ne_zero.mpr hn }, + simp_rw interval_integral.integral_of_le zero_le_one, + refine set_integral_congr measurable_set_Ioc (λ x hx, _), + push_cast, + have hn' : (n : ℂ) ≠ 0, from nat.cast_ne_zero.mpr hn, + have A : (n : ℂ) ^ s = (n : ℂ) ^ (s - 1) * n, + { conv_lhs { rw [(by ring : s = (s - 1) + 1), cpow_add _ _ hn'] }, + simp }, + have B : ((x : ℂ) * ↑n) ^ (s - 1) = (x : ℂ) ^ (s - 1) * ↑n ^ (s - 1), + { rw [←of_real_nat_cast, + mul_cpow_of_real_nonneg hx.1.le (nat.cast_pos.mpr (nat.pos_of_ne_zero hn)).le] }, + rw [A, B, cpow_nat_cast], ring, +end + +/-- The main techical lemma for `Gamma_seq_tendsto_Gamma`, expressing the integral defining the +Gamma function for `0 < re s` as the limit of a sequence of integrals over finite intervals. -/ +lemma approx_Gamma_integral_tendsto_Gamma_integral {s : ℂ} (hs : 0 < re s) : + tendsto (λ n:ℕ, ∫ x:ℝ in 0..n, ↑((1 - x / n) ^ n) * (x:ℂ) ^ (s - 1)) at_top (𝓝 $ Gamma s) := +begin + rw [Gamma_eq_integral hs], + -- We apply dominated convergence to the following function, which we will show is uniformly + -- bounded above by the Gamma integrand `exp (-x) * x ^ (re s - 1)`. + let f : ℕ → ℝ → ℂ := λ n, indicator (Ioc 0 (n:ℝ)) + (λ x:ℝ, ↑((1 - x / n) ^ n) * (x:ℂ) ^ (s - 1)), + -- integrability of f + have f_ible : ∀ (n:ℕ), integrable (f n) (volume.restrict (Ioi 0)), + { intro n, + rw [integrable_indicator_iff (measurable_set_Ioc : measurable_set (Ioc (_:ℝ) _)), + integrable_on, measure.restrict_restrict_of_subset Ioc_subset_Ioi_self, ←integrable_on, + ←interval_integrable_iff_integrable_Ioc_of_le (by positivity : (0:ℝ) ≤ n)], + apply interval_integrable.continuous_on_mul, + { refine interval_integral.interval_integrable_cpow' _, + rwa [sub_re, one_re, ←zero_sub, sub_lt_sub_iff_right] }, + { apply continuous.continuous_on, continuity } }, + -- pointwise limit of f + have f_tends : ∀ x:ℝ, x ∈ Ioi (0:ℝ) → + tendsto (λ n:ℕ, f n x) at_top (𝓝 $ ↑(real.exp (-x)) * (x:ℂ) ^ (s - 1)), + { intros x hx, + apply tendsto.congr', + show ∀ᶠ n:ℕ in at_top, ↑((1 - x / n) ^ n) * (x:ℂ) ^ (s - 1) = f n x, + { refine eventually.mp (eventually_ge_at_top ⌈x⌉₊) (eventually_of_forall (λ n hn, _)), + rw nat.ceil_le at hn, + dsimp only [f], + rw indicator_of_mem, + exact ⟨hx, hn⟩ }, + { simp_rw mul_comm _ (↑x ^ _), + refine (tendsto.comp (continuous_of_real.tendsto _) _).const_mul _, + convert tendsto_one_plus_div_pow_exp (-x), + ext1 n, + rw [neg_div, ←sub_eq_add_neg] } }, + -- let `convert` identify the remaining goals + convert tendsto_integral_of_dominated_convergence _ (λ n, (f_ible n).1) + (real.Gamma_integral_convergent hs) _ + ((ae_restrict_iff' measurable_set_Ioi).mpr (ae_of_all _ f_tends)), + -- limit of f is the integrand we want + { ext1 n, + rw [integral_indicator (measurable_set_Ioc : measurable_set (Ioc (_:ℝ) _)), + interval_integral.integral_of_le (by positivity: 0 ≤ (n:ℝ)), + measure.restrict_restrict_of_subset Ioc_subset_Ioi_self] }, + -- f is uniformly bounded by the Gamma integrand + { intro n, + refine (ae_restrict_iff' measurable_set_Ioi).mpr (ae_of_all _ (λ x hx, _)), + dsimp only [f], + rcases lt_or_le (n:ℝ) x with hxn | hxn, + { rw [indicator_of_not_mem (not_mem_Ioc_of_gt hxn), norm_zero, + mul_nonneg_iff_right_nonneg_of_pos (exp_pos _)], + exact rpow_nonneg_of_nonneg (le_of_lt hx) _ }, + { rw [indicator_of_mem (mem_Ioc.mpr ⟨hx, hxn⟩), norm_mul, complex.norm_eq_abs, + complex.abs_of_nonneg + (pow_nonneg (sub_nonneg.mpr $ div_le_one_of_le hxn $ by positivity) _), + complex.norm_eq_abs, abs_cpow_eq_rpow_re_of_pos hx, sub_re, one_re, + mul_le_mul_right (rpow_pos_of_pos hx _ )], + exact one_sub_div_pow_le_exp_neg hxn } } +end + +/-- Euler's limit formula for the complex Gamma function. -/ +lemma Gamma_seq_tendsto_Gamma (s : ℂ) : + tendsto (Gamma_seq s) at_top (𝓝 $ Gamma s) := +begin + suffices : ∀ m : ℕ, (-↑m < re s) → tendsto (Gamma_seq s) at_top (𝓝 $ Gamma_aux m s), + { rw Gamma, + apply this, + rw neg_lt, + rcases lt_or_le 0 (re s) with hs | hs, + { exact (neg_neg_of_pos hs).trans_le (nat.cast_nonneg _), }, + { refine (nat.lt_floor_add_one _).trans_le _, + rw [sub_eq_neg_add, nat.floor_add_one (neg_nonneg.mpr hs), nat.cast_add_one] } }, + intro m, + induction m with m IH generalizing s, + { -- Base case: `0 < re s`, so Gamma is given by the integral formula + intro hs, + rw [nat.cast_zero, neg_zero] at hs, + rw [←Gamma_eq_Gamma_aux], + { refine tendsto.congr' _ (approx_Gamma_integral_tendsto_Gamma_integral hs), + refine (eventually_ne_at_top 0).mp (eventually_of_forall (λ n hn, _)), + exact (Gamma_seq_eq_approx_Gamma_integral hs hn).symm }, + { rwa [nat.cast_zero, neg_lt_zero] } }, + { -- Induction step: use recurrence formulae in `s` for Gamma and Gamma_seq + intro hs, + rw [nat.cast_succ, neg_add, ←sub_eq_add_neg, sub_lt_iff_lt_add, ←one_re, ←add_re] at hs, + rw Gamma_aux, + have := tendsto.congr' ((eventually_ne_at_top 0).mp (eventually_of_forall (λ n hn, _))) + ((IH _ hs).div_const s), + swap 3, { exact Gamma_seq_add_one_left s hn }, -- doesn't work if inlined? + conv at this in (_ / _ * _) { rw mul_comm }, + rwa [←mul_one (Gamma_aux m (s + 1) / s), tendsto_mul_iff_of_ne_zero _ (one_ne_zero' ℂ)] at this, + simp_rw add_assoc, + exact tendsto_coe_nat_div_add_at_top (1 + s) } +end + +end complex + +end limit_formula + +section gamma_reflection +/-! ## The reflection formula -/ + +open_locale real +namespace complex + +lemma Gamma_seq_mul (z : ℂ) {n : ℕ} (hn : n ≠ 0) : + Gamma_seq z n * Gamma_seq (1 - z) n = + n / (n + 1 - z) * (1 / (z * ∏ j in finset.range n, (1 - z ^ 2 / (j + 1) ^ 2))) := +begin + -- also true for n = 0 but we don't need it + have aux : ∀ (a b c d : ℂ), a * b * (c * d) = a * c * (b * d), by { intros, ring }, + rw [Gamma_seq, Gamma_seq, div_mul_div_comm, aux, ←pow_two], + have : (n : ℂ) ^ z * n ^ (1 - z) = n, + { rw [←cpow_add _ _ (nat.cast_ne_zero.mpr hn), add_sub_cancel'_right, cpow_one] }, + rw [this, finset.prod_range_succ', finset.prod_range_succ, aux, ←finset.prod_mul_distrib, + nat.cast_zero, add_zero, add_comm (1 - z) n, ←add_sub_assoc], + have : ∀ (j : ℕ), (z + ↑(j + 1)) * (1 - z + ↑j) = ↑((j + 1) ^ 2) * (1 - z ^ 2 / (↑j + 1) ^ 2), + { intro j, + push_cast, + have : (j:ℂ) + 1 ≠ 0, by { rw [←nat.cast_succ, nat.cast_ne_zero], exact nat.succ_ne_zero j }, + field_simp, ring }, + simp_rw this, + rw [finset.prod_mul_distrib, ←nat.cast_prod, finset.prod_pow, + finset.prod_range_add_one_eq_factorial, nat.cast_pow, + (by {intros, ring} : ∀ (a b c d : ℂ), a * b * (c * d) = a * (d * (b * c))), + ←div_div, mul_div_cancel, ←div_div, mul_comm z _, mul_one_div], + exact pow_ne_zero 2 (nat.cast_ne_zero.mpr $ nat.factorial_ne_zero n), +end + +/-- Euler's reflection formula for the complex Gamma function. -/ +theorem Gamma_mul_Gamma_one_sub (z : ℂ) : Gamma z * Gamma (1 - z) = π / sin (π * z) := +begin + have pi_ne : (π : ℂ) ≠ 0, from complex.of_real_ne_zero.mpr pi_ne_zero, + by_cases hs : sin (↑π * z) = 0, + { -- first deal with silly case z = integer + rw [hs, div_zero], + rw [←neg_eq_zero, ←complex.sin_neg, ←mul_neg, complex.sin_eq_zero_iff, mul_comm] at hs, + obtain ⟨k, hk⟩ := hs, + rw [mul_eq_mul_right_iff, eq_false_intro (of_real_ne_zero.mpr pi_pos.ne'), or_false, + neg_eq_iff_neg_eq] at hk, + rw ←hk, + cases k, + { rw [int.cast_of_nat, complex.Gamma_neg_nat_eq_zero, zero_mul] }, + { rw [int.cast_neg_succ_of_nat, neg_neg, nat.cast_add, nat.cast_one, add_comm, sub_add_cancel', + complex.Gamma_neg_nat_eq_zero, mul_zero] } }, + refine tendsto_nhds_unique ((Gamma_seq_tendsto_Gamma z).mul (Gamma_seq_tendsto_Gamma $ 1 - z)) _, + have : ↑π / sin (↑π * z) = 1 * (π / sin (π * z)), by rw one_mul, rw this, + refine tendsto.congr' ((eventually_ne_at_top 0).mp + (eventually_of_forall (λ n hn, (Gamma_seq_mul z hn).symm))) (tendsto.mul _ _), + { convert tendsto_coe_nat_div_add_at_top (1 - z), ext1 n, rw add_sub_assoc }, + { have : ↑π / sin (↑π * z) = 1 / (sin (π * z) / π), by field_simp, rw this, + refine tendsto_const_nhds.div _ (div_ne_zero hs pi_ne), + rw [←tendsto_mul_iff_of_ne_zero tendsto_const_nhds pi_ne, div_mul_cancel _ pi_ne], + convert tendsto_euler_sin_prod z, + ext1 n, rw [mul_comm, ←mul_assoc] }, +end + +/-- The Gamma function does not vanish on `ℂ` (except at non-positive integers, where the function +is mathematically undefined and we set it to `0` by convention). -/ +theorem Gamma_ne_zero {s : ℂ} (hs : ∀ m : ℕ, s ≠ -m) : Gamma s ≠ 0 := +begin + by_cases h_im : s.im = 0, + { have : s = ↑s.re, + { conv_lhs { rw ←complex.re_add_im s }, rw [h_im, of_real_zero, zero_mul, add_zero] }, + rw [this, Gamma_of_real, of_real_ne_zero], + refine real.Gamma_ne_zero (λ n, _), + specialize hs n, + contrapose! hs, + rwa [this, ←of_real_nat_cast, ←of_real_neg, of_real_inj] }, + { have : sin (↑π * s) ≠ 0, + { rw complex.sin_ne_zero_iff, + intro k, + apply_fun im, + rw [of_real_mul_im, ←of_real_int_cast, ←of_real_mul, of_real_im], + exact mul_ne_zero real.pi_pos.ne' h_im }, + have A := div_ne_zero (of_real_ne_zero.mpr real.pi_pos.ne') this, + rw [←complex.Gamma_mul_Gamma_one_sub s, mul_ne_zero_iff] at A, + exact A.1 } +end + +lemma Gamma_eq_zero_iff (s : ℂ) : Gamma s = 0 ↔ ∃ m : ℕ, s = -m := +begin + split, + { contrapose!, exact Gamma_ne_zero }, + { rintro ⟨m, rfl⟩, exact Gamma_neg_nat_eq_zero m }, +end + +end complex + +namespace real + +/-- The sequence with `n`-th term `n ^ s * n! / (s * (s + 1) * ... * (s + n))`, for real `s`. We +will show that this tends to `Γ(s)` as `n → ∞`. -/ +noncomputable def Gamma_seq (s : ℝ) (n : ℕ) := +(n : ℝ) ^ s * n! / ∏ (j : ℕ) in finset.range (n + 1), (s + j) + +/-- Euler's limit formula for the real Gamma function. -/ +lemma Gamma_seq_tendsto_Gamma (s : ℝ) : tendsto (Gamma_seq s) at_top (𝓝 $ Gamma s) := +begin + suffices : tendsto (coe ∘ Gamma_seq s : ℕ → ℂ) at_top (𝓝 $ complex.Gamma s), + from (complex.continuous_re.tendsto (complex.Gamma ↑s)).comp this, + convert complex.Gamma_seq_tendsto_Gamma s, + ext1 n, + dsimp only [Gamma_seq, function.comp_app, complex.Gamma_seq], + push_cast, + rw [complex.of_real_cpow n.cast_nonneg, complex.of_real_nat_cast] +end + +/-- Euler's reflection formula for the real Gamma function. -/ +lemma Gamma_mul_Gamma_one_sub (s : ℝ) : Gamma s * Gamma (1 - s) = π / sin (π * s) := +begin + simp_rw [←complex.of_real_inj, complex.of_real_div, complex.of_real_sin, + complex.of_real_mul, ←complex.Gamma_of_real, complex.of_real_sub, complex.of_real_one], + exact complex.Gamma_mul_Gamma_one_sub s +end + +end real + +end gamma_reflection diff --git a/src/analysis/specific_limits/basic.lean b/src/analysis/specific_limits/basic.lean index 2ec78fcb19695..6213b01d8c5d4 100644 --- a/src/analysis/specific_limits/basic.lean +++ b/src/analysis/specific_limits/basic.lean @@ -7,6 +7,7 @@ import algebra.geom_sum import order.filter.archimedean import order.iterate import topology.instances.ennreal +import topology.algebra.algebra /-! # A collection of specific limit computations @@ -41,6 +42,34 @@ lemma tendsto_one_div_add_at_top_nhds_0_nat : suffices tendsto (λ n : ℕ, 1 / (↑(n + 1) : ℝ)) at_top (𝓝 0), by simpa, (tendsto_add_at_top_iff_nat 1).2 (tendsto_const_div_at_top_nhds_0_nat 1) +/-- The limit of `n / (n + x)` is 1, for any constant `x` (valid in `ℝ` or any topological division +algebra over `ℝ`, e.g., `ℂ`). + +TODO: introduce a typeclass saying that `1 / n` tends to 0 at top, making it possible to get this +statement simultaneously on `ℚ`, `ℝ` and `ℂ`. -/ +lemma tendsto_coe_nat_div_add_at_top + {𝕜 : Type*} [division_ring 𝕜] [topological_space 𝕜] [char_zero 𝕜] [algebra ℝ 𝕜] + [has_continuous_smul ℝ 𝕜] [topological_division_ring 𝕜] + (x : 𝕜) : + tendsto (λ n:ℕ, (n:𝕜) / (n + x)) at_top (𝓝 1) := +begin + refine tendsto.congr' ((eventually_ne_at_top 0).mp (eventually_of_forall (λ n hn, _))) _, + { exact λ n:ℕ, 1 / (1 + x / n) }, + { field_simp [nat.cast_ne_zero.mpr hn] }, + { have : 𝓝 (1:𝕜) = 𝓝 (1 / (1 + x * ↑(0:ℝ))), + by rw [algebra_map.coe_zero, mul_zero, add_zero, div_one], + rw this, + refine tendsto_const_nhds.div (tendsto_const_nhds.add _) (by simp), + simp_rw div_eq_mul_inv, + refine (tendsto_const_nhds.mul _), + have : (λ n : ℕ, (n : 𝕜)⁻¹) = (λ n : ℕ, ↑((n : ℝ)⁻¹)), + { ext1 n, + rw [←(map_nat_cast (algebra_map ℝ 𝕜) n), ←map_inv₀ (algebra_map ℝ 𝕜)], + refl, }, + rw this, + exact ((continuous_algebra_map ℝ 𝕜).tendsto _).comp tendsto_inverse_at_top_nhds_0_nat } +end + /-! ### Powers -/ lemma tendsto_add_one_pow_at_top_at_top_of_pos [linear_ordered_semiring α] [archimedean α] {r : α} diff --git a/src/data/complex/exponential.lean b/src/data/complex/exponential.lean index f41c5375fa3cc..4aebd98bd083d 100644 --- a/src/data/complex/exponential.lean +++ b/src/data/complex/exponential.lean @@ -1625,6 +1625,18 @@ begin exact add_one_le_exp_of_nonpos h.le, end +lemma one_sub_div_pow_le_exp_neg {n : ℕ} {t : ℝ} (ht' : t ≤ n) : (1 - t / n) ^ n ≤ exp (-t) := +begin + rcases eq_or_ne n 0 with rfl | hn, + { simp, rwa nat.cast_zero at ht' }, + convert pow_le_pow_of_le_left _ (add_one_le_exp (-(t / n))) n, + { abel }, + { rw ←real.exp_nat_mul, congr' 1, + field_simp [nat.cast_ne_zero.mpr hn], ring }, + { rwa [add_comm, ←sub_eq_add_neg, sub_nonneg, div_le_one], + positivity } +end + end real namespace tactic From f1e061e3caef3022f0daa99d670ecf2c30e0b5c6 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 15 Feb 2023 11:30:37 +0000 Subject: [PATCH 0480/1291] chore(set_theory/ordinal/arithmetic): move power and logarithm to a new file (#18441) This represents a significant fraction of the number of lines. The lemmas are moved without any changes. --- src/set_theory/ordinal/arithmetic.lean | 412 +---------------- .../ordinal/cantor_normal_form.lean | 1 + src/set_theory/ordinal/exponential.lean | 433 ++++++++++++++++++ src/set_theory/ordinal/fixed_point.lean | 1 + 4 files changed, 437 insertions(+), 410 deletions(-) create mode 100644 src/set_theory/ordinal/exponential.lean diff --git a/src/set_theory/ordinal/arithmetic.lean b/src/set_theory/ordinal/arithmetic.lean index 5f8e3b0edbf3c..7a3bd7d4615cf 100644 --- a/src/set_theory/ordinal/arithmetic.lean +++ b/src/set_theory/ordinal/arithmetic.lean @@ -28,8 +28,8 @@ successor ordinals and limit ordinals, in `limit_rec_on`. * `order.succ o = o + 1` is the successor of `o`. * `pred o` if the predecessor of `o`. If `o` is not a successor, we set `pred o = o`. -We also define the power function and the logarithm function on ordinals, and discuss the properties -of casts of natural numbers of and of `ω` with respect to these operations. +We discuss the properties of casts of natural numbers of and of `ω` with respect to these +operations. Some properties of the operations are also used to discuss general tools on ordinals: @@ -1740,383 +1740,6 @@ end end -/-! ### Ordinal exponential -/ - -/-- The ordinal exponential, defined by transfinite recursion. -/ -instance : has_pow ordinal ordinal := -⟨λ a b, if a = 0 then 1 - b else limit_rec_on b 1 (λ _ IH, IH * a) (λ b _, bsup.{u u} b)⟩ - -local infixr (name := ordinal.pow) ^ := @pow ordinal ordinal ordinal.has_pow - -theorem opow_def (a b : ordinal) : - a ^ b = if a = 0 then 1 - b else limit_rec_on b 1 (λ _ IH, IH * a) (λ b _, bsup.{u u} b) := -rfl - -theorem zero_opow' (a : ordinal) : 0 ^ a = 1 - a := -by simp only [opow_def, if_pos rfl] - -@[simp] theorem zero_opow {a : ordinal} (a0 : a ≠ 0) : 0 ^ a = 0 := -by rwa [zero_opow', ordinal.sub_eq_zero_iff_le, one_le_iff_ne_zero] - -@[simp] theorem opow_zero (a : ordinal) : a ^ 0 = 1 := -by by_cases a = 0; [simp only [opow_def, if_pos h, sub_zero], -simp only [opow_def, if_neg h, limit_rec_on_zero]] - -@[simp] theorem opow_succ (a b : ordinal) : a ^ succ b = a ^ b * a := -if h : a = 0 then by subst a; simp only [zero_opow (succ_ne_zero _), mul_zero] -else by simp only [opow_def, limit_rec_on_succ, if_neg h] - -theorem opow_limit {a b : ordinal} (a0 : a ≠ 0) (h : is_limit b) : - a ^ b = bsup.{u u} b (λ c _, a ^ c) := -by simp only [opow_def, if_neg a0]; rw limit_rec_on_limit _ _ _ _ h; refl - -theorem opow_le_of_limit {a b c : ordinal} (a0 : a ≠ 0) (h : is_limit b) : - a ^ b ≤ c ↔ ∀ b' < b, a ^ b' ≤ c := -by rw [opow_limit a0 h, bsup_le_iff] - -theorem lt_opow_of_limit {a b c : ordinal} (b0 : b ≠ 0) (h : is_limit c) : - a < b ^ c ↔ ∃ c' < c, a < b ^ c' := -by rw [← not_iff_not, not_exists]; simp only [not_lt, opow_le_of_limit b0 h, exists_prop, not_and] - -@[simp] theorem opow_one (a : ordinal) : a ^ 1 = a := -by rw [← succ_zero, opow_succ]; simp only [opow_zero, one_mul] - -@[simp] theorem one_opow (a : ordinal) : 1 ^ a = 1 := -begin - apply limit_rec_on a, - { simp only [opow_zero] }, - { intros _ ih, simp only [opow_succ, ih, mul_one] }, - refine λ b l IH, eq_of_forall_ge_iff (λ c, _), - rw [opow_le_of_limit ordinal.one_ne_zero l], - exact ⟨λ H, by simpa only [opow_zero] using H 0 l.pos, - λ H b' h, by rwa IH _ h⟩, -end - -theorem opow_pos {a : ordinal} (b) - (a0 : 0 < a) : 0 < a ^ b := -begin - have h0 : 0 < a ^ 0, {simp only [opow_zero, zero_lt_one]}, - apply limit_rec_on b, - { exact h0 }, - { intros b IH, rw [opow_succ], - exact mul_pos IH a0 }, - { exact λ b l _, (lt_opow_of_limit (ordinal.pos_iff_ne_zero.1 a0) l).2 - ⟨0, l.pos, h0⟩ }, -end - -theorem opow_ne_zero {a : ordinal} (b) - (a0 : a ≠ 0) : a ^ b ≠ 0 := -ordinal.pos_iff_ne_zero.1 $ opow_pos b $ ordinal.pos_iff_ne_zero.2 a0 - -theorem opow_is_normal {a : ordinal} (h : 1 < a) : is_normal ((^) a) := -have a0 : 0 < a, from zero_lt_one.trans h, -⟨λ b, by simpa only [mul_one, opow_succ] using - (mul_lt_mul_iff_left (opow_pos b a0)).2 h, - λ b l c, opow_le_of_limit (ne_of_gt a0) l⟩ - -theorem opow_lt_opow_iff_right {a b c : ordinal} - (a1 : 1 < a) : a ^ b < a ^ c ↔ b < c := -(opow_is_normal a1).lt_iff - -theorem opow_le_opow_iff_right {a b c : ordinal} - (a1 : 1 < a) : a ^ b ≤ a ^ c ↔ b ≤ c := -(opow_is_normal a1).le_iff - -theorem opow_right_inj {a b c : ordinal} - (a1 : 1 < a) : a ^ b = a ^ c ↔ b = c := -(opow_is_normal a1).inj - -theorem opow_is_limit {a b : ordinal} - (a1 : 1 < a) : is_limit b → is_limit (a ^ b) := -(opow_is_normal a1).is_limit - -theorem opow_is_limit_left {a b : ordinal} - (l : is_limit a) (hb : b ≠ 0) : is_limit (a ^ b) := -begin - rcases zero_or_succ_or_limit b with e|⟨b,rfl⟩|l', - { exact absurd e hb }, - { rw opow_succ, - exact mul_is_limit (opow_pos _ l.pos) l }, - { exact opow_is_limit l.one_lt l' } -end - -theorem opow_le_opow_right {a b c : ordinal} - (h₁ : 0 < a) (h₂ : b ≤ c) : a ^ b ≤ a ^ c := -begin - cases lt_or_eq_of_le (one_le_iff_pos.2 h₁) with h₁ h₁, - { exact (opow_le_opow_iff_right h₁).2 h₂ }, - { subst a, simp only [one_opow] } -end - -theorem opow_le_opow_left {a b : ordinal} (c) - (ab : a ≤ b) : a ^ c ≤ b ^ c := -begin - by_cases a0 : a = 0, - { subst a, by_cases c0 : c = 0, - { subst c, simp only [opow_zero] }, - { simp only [zero_opow c0, ordinal.zero_le] } }, - { apply limit_rec_on c, - { simp only [opow_zero] }, - { intros c IH, simpa only [opow_succ] using mul_le_mul' IH ab }, - { exact λ c l IH, (opow_le_of_limit a0 l).2 - (λ b' h, (IH _ h).trans (opow_le_opow_right - ((ordinal.pos_iff_ne_zero.2 a0).trans_le ab) h.le)) } } -end - -theorem left_le_opow (a : ordinal) {b : ordinal} (b1 : 0 < b) : a ≤ a ^ b := -begin - nth_rewrite 0 ←opow_one a, - cases le_or_gt a 1 with a1 a1, - { cases lt_or_eq_of_le a1 with a0 a1, - { rw lt_one_iff_zero at a0, - rw [a0, zero_opow ordinal.one_ne_zero], - exact ordinal.zero_le _ }, - rw [a1, one_opow, one_opow] }, - rwa [opow_le_opow_iff_right a1, one_le_iff_pos] -end - -theorem right_le_opow {a : ordinal} (b) (a1 : 1 < a) : b ≤ a ^ b := -(opow_is_normal a1).self_le _ - -theorem opow_lt_opow_left_of_succ {a b c : ordinal} - (ab : a < b) : a ^ succ c < b ^ succ c := -by { rw [opow_succ, opow_succ], exact - (mul_le_mul_right' (opow_le_opow_left c ab.le) a).trans_lt - (mul_lt_mul_of_pos_left ab (opow_pos c ((ordinal.zero_le a).trans_lt ab))) } - -theorem opow_add (a b c : ordinal) : a ^ (b + c) = a ^ b * a ^ c := -begin - rcases eq_or_ne a 0 with rfl | a0, - { rcases eq_or_ne c 0 with rfl | c0, { simp }, - have : b + c ≠ 0 := ((ordinal.pos_iff_ne_zero.2 c0).trans_le (le_add_left _ _)).ne', - simp only [zero_opow c0, zero_opow this, mul_zero] }, - rcases eq_or_lt_of_le (one_le_iff_ne_zero.2 a0) with rfl | a1, - { simp only [one_opow, mul_one] }, - apply limit_rec_on c, - { simp }, - { intros c IH, - rw [add_succ, opow_succ, IH, opow_succ, mul_assoc] }, - { intros c l IH, - refine eq_of_forall_ge_iff (λ d, (((opow_is_normal a1).trans - (add_is_normal b)).limit_le l).trans _), - dsimp only [function.comp], - simp only [IH] {contextual := tt}, - exact (((mul_is_normal $ opow_pos b (ordinal.pos_iff_ne_zero.2 a0)).trans - (opow_is_normal a1)).limit_le l).symm } -end - -theorem opow_one_add (a b : ordinal) : a ^ (1 + b) = a * a ^ b := -by rw [opow_add, opow_one] - -theorem opow_dvd_opow (a) {b c : ordinal} - (h : b ≤ c) : a ^ b ∣ a ^ c := -by { rw [← ordinal.add_sub_cancel_of_le h, opow_add], apply dvd_mul_right } - -theorem opow_dvd_opow_iff {a b c : ordinal} - (a1 : 1 < a) : a ^ b ∣ a ^ c ↔ b ≤ c := -⟨λ h, le_of_not_lt $ λ hn, - not_le_of_lt ((opow_lt_opow_iff_right a1).2 hn) $ - le_of_dvd (opow_ne_zero _ $ one_le_iff_ne_zero.1 $ a1.le) h, -opow_dvd_opow _⟩ - -theorem opow_mul (a b c : ordinal) : a ^ (b * c) = (a ^ b) ^ c := -begin - by_cases b0 : b = 0, {simp only [b0, zero_mul, opow_zero, one_opow]}, - by_cases a0 : a = 0, - { subst a, - by_cases c0 : c = 0, {simp only [c0, mul_zero, opow_zero]}, - simp only [zero_opow b0, zero_opow c0, zero_opow (mul_ne_zero b0 c0)] }, - cases eq_or_lt_of_le (one_le_iff_ne_zero.2 a0) with a1 a1, - { subst a1, simp only [one_opow] }, - apply limit_rec_on c, - { simp only [mul_zero, opow_zero] }, - { intros c IH, - rw [mul_succ, opow_add, IH, opow_succ] }, - { intros c l IH, - refine eq_of_forall_ge_iff (λ d, (((opow_is_normal a1).trans - (mul_is_normal (ordinal.pos_iff_ne_zero.2 b0))).limit_le l).trans _), - dsimp only [function.comp], - simp only [IH] {contextual := tt}, - exact (opow_le_of_limit (opow_ne_zero _ a0) l).symm } -end - -/-! ### Ordinal logarithm -/ - -/-- The ordinal logarithm is the solution `u` to the equation `x = b ^ u * v + w` where `v < b` and - `w < b ^ u`. -/ -@[pp_nodot] def log (b : ordinal) (x : ordinal) : ordinal := -if h : 1 < b then pred (Inf {o | x < b ^ o}) else 0 - -/-- The set in the definition of `log` is nonempty. -/ -theorem log_nonempty {b x : ordinal} (h : 1 < b) : {o | x < b ^ o}.nonempty := -⟨_, succ_le_iff.1 (right_le_opow _ h)⟩ - -theorem log_def {b : ordinal} (h : 1 < b) (x : ordinal) : log b x = pred (Inf {o | x < b ^ o}) := -by simp only [log, dif_pos h] - -theorem log_of_not_one_lt_left {b : ordinal} (h : ¬ 1 < b) (x : ordinal) : log b x = 0 := -by simp only [log, dif_neg h] - -theorem log_of_left_le_one {b : ordinal} (h : b ≤ 1) : ∀ x, log b x = 0 := -log_of_not_one_lt_left h.not_lt - -@[simp] lemma log_zero_left : ∀ b, log 0 b = 0 := -log_of_left_le_one zero_le_one - -@[simp] theorem log_zero_right (b : ordinal) : log b 0 = 0 := -if b1 : 1 < b then begin - rw [log_def b1, ← ordinal.le_zero, pred_le], - apply cInf_le', - dsimp, - rw [succ_zero, opow_one], - exact zero_lt_one.trans b1 -end -else by simp only [log_of_not_one_lt_left b1] - -@[simp] theorem log_one_left : ∀ b, log 1 b = 0 := -log_of_left_le_one le_rfl - -theorem succ_log_def {b x : ordinal} (hb : 1 < b) (hx : x ≠ 0) : - succ (log b x) = Inf {o | x < b ^ o} := -begin - let t := Inf {o | x < b ^ o}, - have : x < b ^ t := Inf_mem (log_nonempty hb), - rcases zero_or_succ_or_limit t with h|h|h, - { refine ((one_le_iff_ne_zero.2 hx).not_lt _).elim, - simpa only [h, opow_zero] }, - { rw [show log b x = pred t, from log_def hb x, - succ_pred_iff_is_succ.2 h] }, - { rcases (lt_opow_of_limit (zero_lt_one.trans hb).ne' h).1 this with ⟨a, h₁, h₂⟩, - exact h₁.not_le.elim ((le_cInf_iff'' (log_nonempty hb)).1 le_rfl a h₂) } -end - -theorem lt_opow_succ_log_self {b : ordinal} (hb : 1 < b) (x : ordinal) : x < b ^ succ (log b x) := -begin - rcases eq_or_ne x 0 with rfl | hx, - { apply opow_pos _ (zero_lt_one.trans hb) }, - { rw succ_log_def hb hx, - exact Inf_mem (log_nonempty hb) } -end - -theorem opow_log_le_self (b) {x : ordinal} (hx : x ≠ 0) : b ^ log b x ≤ x := -begin - rcases eq_or_ne b 0 with rfl | b0, - { rw zero_opow', - refine (sub_le_self _ _).trans (one_le_iff_ne_zero.2 hx) }, - rcases lt_or_eq_of_le (one_le_iff_ne_zero.2 b0) with hb | rfl, - { refine le_of_not_lt (λ h, (lt_succ (log b x)).not_le _), - have := @cInf_le' _ _ {o | x < b ^ o} _ h, - rwa ←succ_log_def hb hx at this }, - { rwa [one_opow, one_le_iff_ne_zero] } -end - -/-- `opow b` and `log b` (almost) form a Galois connection. -/ -theorem opow_le_iff_le_log {b x c : ordinal} (hb : 1 < b) (hx : x ≠ 0) : b ^ c ≤ x ↔ c ≤ log b x := -⟨λ h, le_of_not_lt $ λ hn, - (lt_opow_succ_log_self hb x).not_le $ - ((opow_le_opow_iff_right hb).2 (succ_le_of_lt hn)).trans h, -λ h, ((opow_le_opow_iff_right hb).2 h).trans (opow_log_le_self b hx)⟩ - -theorem lt_opow_iff_log_lt {b x c : ordinal} (hb : 1 < b) (hx : x ≠ 0) : x < b ^ c ↔ log b x < c := -lt_iff_lt_of_le_iff_le (opow_le_iff_le_log hb hx) - -theorem log_pos {b o : ordinal} (hb : 1 < b) (ho : o ≠ 0) (hbo : b ≤ o) : 0 < log b o := -by rwa [←succ_le_iff, succ_zero, ←opow_le_iff_le_log hb ho, opow_one] - -theorem log_eq_zero {b o : ordinal} (hbo : o < b) : log b o = 0 := -begin - rcases eq_or_ne o 0 with rfl | ho, - { exact log_zero_right b }, - cases le_or_lt b 1 with hb hb, - { rcases le_one_iff.1 hb with rfl | rfl, - { exact log_zero_left o }, - { exact log_one_left o } }, - { rwa [←ordinal.le_zero, ←lt_succ_iff, succ_zero, ←lt_opow_iff_log_lt hb ho, opow_one] } -end - -@[mono] theorem log_mono_right (b) {x y : ordinal} (xy : x ≤ y) : log b x ≤ log b y := -if hx : x = 0 then by simp only [hx, log_zero_right, ordinal.zero_le] else -if hb : 1 < b then - (opow_le_iff_le_log hb (lt_of_lt_of_le (ordinal.pos_iff_ne_zero.2 hx) xy).ne').1 $ - (opow_log_le_self _ hx).trans xy -else by simp only [log_of_not_one_lt_left hb, ordinal.zero_le] - -theorem log_le_self (b x : ordinal) : log b x ≤ x := -if hx : x = 0 then by simp only [hx, log_zero_right, ordinal.zero_le] else -if hb : 1 < b then (right_le_opow _ hb).trans (opow_log_le_self b hx) -else by simp only [log_of_not_one_lt_left hb, ordinal.zero_le] - -@[simp] theorem log_one_right (b : ordinal) : log b 1 = 0 := -if hb : 1 < b then log_eq_zero hb else log_of_not_one_lt_left hb 1 - -theorem mod_opow_log_lt_self (b : ordinal) {o : ordinal} (ho : o ≠ 0) : o % b ^ log b o < o := -begin - rcases eq_or_ne b 0 with rfl | hb, - { simpa using ordinal.pos_iff_ne_zero.2 ho }, - { exact (mod_lt _ $ opow_ne_zero _ hb).trans_le (opow_log_le_self _ ho) } -end - -theorem log_mod_opow_log_lt_log_self {b o : ordinal} (hb : 1 < b) (ho : o ≠ 0) (hbo : b ≤ o) : - log b (o % b ^ log b o) < log b o := -begin - cases eq_or_ne (o % b ^ log b o) 0, - { rw [h, log_zero_right], - apply log_pos hb ho hbo }, - { rw [←succ_le_iff, succ_log_def hb h], - apply cInf_le', - apply mod_lt, - rw ←ordinal.pos_iff_ne_zero, - exact opow_pos _ (zero_lt_one.trans hb) } -end - -lemma opow_mul_add_pos {b v : ordinal} (hb : b ≠ 0) (u) (hv : v ≠ 0) (w) : 0 < b ^ u * v + w := -(opow_pos u $ ordinal.pos_iff_ne_zero.2 hb).trans_le $ - (le_mul_left _ $ ordinal.pos_iff_ne_zero.2 hv).trans $ le_add_right _ _ - -lemma opow_mul_add_lt_opow_mul_succ {b u w : ordinal} (v : ordinal) (hw : w < b ^ u) : - b ^ u * v + w < b ^ u * (succ v) := -by rwa [mul_succ, add_lt_add_iff_left] - -lemma opow_mul_add_lt_opow_succ {b u v w : ordinal} (hvb : v < b) (hw : w < b ^ u) : - b ^ u * v + w < b ^ (succ u) := -begin - convert (opow_mul_add_lt_opow_mul_succ v hw).trans_le (mul_le_mul_left' (succ_le_of_lt hvb) _), - exact opow_succ b u -end - -theorem log_opow_mul_add {b u v w : ordinal} (hb : 1 < b) (hv : v ≠ 0) (hvb : v < b) - (hw : w < b ^ u) : log b (b ^ u * v + w) = u := -begin - have hne' := (opow_mul_add_pos (zero_lt_one.trans hb).ne' u hv w).ne', - by_contra' hne, - cases lt_or_gt_of_ne hne with h h, - { rw ←lt_opow_iff_log_lt hb hne' at h, - exact h.not_le ((le_mul_left _ (ordinal.pos_iff_ne_zero.2 hv)).trans (le_add_right _ _)) }, - { change _ < _ at h, - rw [←succ_le_iff, ←opow_le_iff_le_log hb hne'] at h, - exact (not_lt_of_le h) (opow_mul_add_lt_opow_succ hvb hw) } -end - -theorem log_opow {b : ordinal} (hb : 1 < b) (x : ordinal) : log b (b ^ x) = x := -begin - convert log_opow_mul_add hb zero_ne_one.symm hb (opow_pos x (zero_lt_one.trans hb)), - rw [add_zero, mul_one] -end - -theorem div_opow_log_lt {b : ordinal} (o : ordinal) (hb : 1 < b) : o / b ^ log b o < b := -begin - rw [div_lt (opow_pos _ (zero_lt_one.trans hb)).ne', ←opow_succ], - exact lt_opow_succ_log_self hb o -end - -theorem add_log_le_log_mul {x y : ordinal} (b : ordinal) (hx : x ≠ 0) (hy : y ≠ 0) : - log b x + log b y ≤ log b (x * y) := -begin - by_cases hb : 1 < b, - { rw [←opow_le_iff_le_log hb (mul_ne_zero hx hy), opow_add], - exact mul_le_mul' (opow_log_le_self b hx) (opow_log_le_self b hy) }, - simp only [log_of_not_one_lt_left hb, zero_add] -end /-! ### Casting naturals into ordinals, compatibility with operations -/ @@ -2127,10 +1750,6 @@ by { rw [←nat.cast_one, ←nat.cast_add, add_comm], refl } | 0 := by simp | (n+1) := by rw [nat.mul_succ, nat.cast_add, nat_cast_mul, nat.cast_succ, mul_add_one] -@[simp, norm_cast] theorem nat_cast_opow (m : ℕ) : ∀ n : ℕ, ((pow m n : ℕ) : ordinal) = m ^ n -| 0 := by simp -| (n+1) := by rw [pow_succ', nat_cast_mul, nat_cast_opow, nat.cast_succ, add_one_eq_succ, opow_succ] - @[simp, norm_cast] theorem nat_cast_le {m n : ℕ} : (m : ordinal) ≤ n ↔ m ≤ n := by rw [←cardinal.ord_nat, ←cardinal.ord_nat, cardinal.ord_le_ord, cardinal.nat_cast_le] @@ -2310,35 +1929,8 @@ begin { exact (mul_is_normal ho).apply_omega } end -local infixr (name := ordinal.pow) ^ := @pow ordinal ordinal ordinal.has_pow -theorem sup_opow_nat {o : ordinal} (ho : 0 < o) : sup (λ n : ℕ, o ^ n) = o ^ ω := -begin - rcases lt_or_eq_of_le (one_le_iff_pos.2 ho) with ho₁ | rfl, - { exact (opow_is_normal ho₁).apply_omega }, - { rw one_opow, - refine le_antisymm (sup_le (λ n, by rw one_opow)) _, - convert le_sup _ 0, - rw [nat.cast_zero, opow_zero] } -end - end ordinal -namespace tactic -open ordinal positivity - -/-- Extension for the `positivity` tactic: `ordinal.opow` takes positive values on positive inputs. --/ -@[positivity] -meta def positivity_opow : expr → tactic strictness -| `(@has_pow.pow _ _ %%inst %%a %%b) := do - strictness_a ← core a, - match strictness_a with - | positive p := positive <$> mk_app ``opow_pos [b, p] - | _ := failed -- We already know that `0 ≤ x` for all `x : ordinal` - end -| _ := failed - -end tactic variables {α : Type u} {r : α → α → Prop} {a b : α} diff --git a/src/set_theory/ordinal/cantor_normal_form.lean b/src/set_theory/ordinal/cantor_normal_form.lean index 148bc847319ef..0447f5158c919 100644 --- a/src/set_theory/ordinal/cantor_normal_form.lean +++ b/src/set_theory/ordinal/cantor_normal_form.lean @@ -5,6 +5,7 @@ Authors: Mario Carneiro -/ import set_theory.ordinal.arithmetic +import set_theory.ordinal.exponential /-! # Cantor Normal Form diff --git a/src/set_theory/ordinal/exponential.lean b/src/set_theory/ordinal/exponential.lean new file mode 100644 index 0000000000000..ac04319279e5e --- /dev/null +++ b/src/set_theory/ordinal/exponential.lean @@ -0,0 +1,433 @@ +/- +Copyright (c) 2017 Mario Carneiro. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Mario Carneiro, Floris van Doorn, Violeta Hernández Palacios +-/ +import set_theory.ordinal.arithmetic + +/-! # Ordinal exponential + +In this file we define the power function and the logarithm function on ordinals, + +-/ + +noncomputable theory + +open function cardinal set equiv order +open_locale classical cardinal ordinal + +universes u v w + +namespace ordinal + +/-- The ordinal exponential, defined by transfinite recursion. -/ +instance : has_pow ordinal ordinal := +⟨λ a b, if a = 0 then 1 - b else limit_rec_on b 1 (λ _ IH, IH * a) (λ b _, bsup.{u u} b)⟩ + +local infixr (name := ordinal.pow) ^ := @pow ordinal ordinal ordinal.has_pow + +theorem opow_def (a b : ordinal) : + a ^ b = if a = 0 then 1 - b else limit_rec_on b 1 (λ _ IH, IH * a) (λ b _, bsup.{u u} b) := +rfl + +theorem zero_opow' (a : ordinal) : 0 ^ a = 1 - a := +by simp only [opow_def, if_pos rfl] + +@[simp] theorem zero_opow {a : ordinal} (a0 : a ≠ 0) : 0 ^ a = 0 := +by rwa [zero_opow', ordinal.sub_eq_zero_iff_le, one_le_iff_ne_zero] + +@[simp] theorem opow_zero (a : ordinal) : a ^ 0 = 1 := +by by_cases a = 0; [simp only [opow_def, if_pos h, sub_zero], +simp only [opow_def, if_neg h, limit_rec_on_zero]] + +@[simp] theorem opow_succ (a b : ordinal) : a ^ succ b = a ^ b * a := +if h : a = 0 then by subst a; simp only [zero_opow (succ_ne_zero _), mul_zero] +else by simp only [opow_def, limit_rec_on_succ, if_neg h] + +theorem opow_limit {a b : ordinal} (a0 : a ≠ 0) (h : is_limit b) : + a ^ b = bsup.{u u} b (λ c _, a ^ c) := +by simp only [opow_def, if_neg a0]; rw limit_rec_on_limit _ _ _ _ h; refl + +theorem opow_le_of_limit {a b c : ordinal} (a0 : a ≠ 0) (h : is_limit b) : + a ^ b ≤ c ↔ ∀ b' < b, a ^ b' ≤ c := +by rw [opow_limit a0 h, bsup_le_iff] + +theorem lt_opow_of_limit {a b c : ordinal} (b0 : b ≠ 0) (h : is_limit c) : + a < b ^ c ↔ ∃ c' < c, a < b ^ c' := +by rw [← not_iff_not, not_exists]; simp only [not_lt, opow_le_of_limit b0 h, exists_prop, not_and] + +@[simp] theorem opow_one (a : ordinal) : a ^ 1 = a := +by rw [← succ_zero, opow_succ]; simp only [opow_zero, one_mul] + +@[simp] theorem one_opow (a : ordinal) : 1 ^ a = 1 := +begin + apply limit_rec_on a, + { simp only [opow_zero] }, + { intros _ ih, simp only [opow_succ, ih, mul_one] }, + refine λ b l IH, eq_of_forall_ge_iff (λ c, _), + rw [opow_le_of_limit ordinal.one_ne_zero l], + exact ⟨λ H, by simpa only [opow_zero] using H 0 l.pos, + λ H b' h, by rwa IH _ h⟩, +end + +theorem opow_pos {a : ordinal} (b) + (a0 : 0 < a) : 0 < a ^ b := +begin + have h0 : 0 < a ^ 0, {simp only [opow_zero, zero_lt_one]}, + apply limit_rec_on b, + { exact h0 }, + { intros b IH, rw [opow_succ], + exact mul_pos IH a0 }, + { exact λ b l _, (lt_opow_of_limit (ordinal.pos_iff_ne_zero.1 a0) l).2 + ⟨0, l.pos, h0⟩ }, +end + +theorem opow_ne_zero {a : ordinal} (b) + (a0 : a ≠ 0) : a ^ b ≠ 0 := +ordinal.pos_iff_ne_zero.1 $ opow_pos b $ ordinal.pos_iff_ne_zero.2 a0 + +theorem opow_is_normal {a : ordinal} (h : 1 < a) : is_normal ((^) a) := +have a0 : 0 < a, from zero_lt_one.trans h, +⟨λ b, by simpa only [mul_one, opow_succ] using + (mul_lt_mul_iff_left (opow_pos b a0)).2 h, + λ b l c, opow_le_of_limit (ne_of_gt a0) l⟩ + +theorem opow_lt_opow_iff_right {a b c : ordinal} + (a1 : 1 < a) : a ^ b < a ^ c ↔ b < c := +(opow_is_normal a1).lt_iff + +theorem opow_le_opow_iff_right {a b c : ordinal} + (a1 : 1 < a) : a ^ b ≤ a ^ c ↔ b ≤ c := +(opow_is_normal a1).le_iff + +theorem opow_right_inj {a b c : ordinal} + (a1 : 1 < a) : a ^ b = a ^ c ↔ b = c := +(opow_is_normal a1).inj + +theorem opow_is_limit {a b : ordinal} + (a1 : 1 < a) : is_limit b → is_limit (a ^ b) := +(opow_is_normal a1).is_limit + +theorem opow_is_limit_left {a b : ordinal} + (l : is_limit a) (hb : b ≠ 0) : is_limit (a ^ b) := +begin + rcases zero_or_succ_or_limit b with e|⟨b,rfl⟩|l', + { exact absurd e hb }, + { rw opow_succ, + exact mul_is_limit (opow_pos _ l.pos) l }, + { exact opow_is_limit l.one_lt l' } +end + +theorem opow_le_opow_right {a b c : ordinal} + (h₁ : 0 < a) (h₂ : b ≤ c) : a ^ b ≤ a ^ c := +begin + cases lt_or_eq_of_le (one_le_iff_pos.2 h₁) with h₁ h₁, + { exact (opow_le_opow_iff_right h₁).2 h₂ }, + { subst a, simp only [one_opow] } +end + +theorem opow_le_opow_left {a b : ordinal} (c) + (ab : a ≤ b) : a ^ c ≤ b ^ c := +begin + by_cases a0 : a = 0, + { subst a, by_cases c0 : c = 0, + { subst c, simp only [opow_zero] }, + { simp only [zero_opow c0, ordinal.zero_le] } }, + { apply limit_rec_on c, + { simp only [opow_zero] }, + { intros c IH, simpa only [opow_succ] using mul_le_mul' IH ab }, + { exact λ c l IH, (opow_le_of_limit a0 l).2 + (λ b' h, (IH _ h).trans (opow_le_opow_right + ((ordinal.pos_iff_ne_zero.2 a0).trans_le ab) h.le)) } } +end + +theorem left_le_opow (a : ordinal) {b : ordinal} (b1 : 0 < b) : a ≤ a ^ b := +begin + nth_rewrite 0 ←opow_one a, + cases le_or_gt a 1 with a1 a1, + { cases lt_or_eq_of_le a1 with a0 a1, + { rw lt_one_iff_zero at a0, + rw [a0, zero_opow ordinal.one_ne_zero], + exact ordinal.zero_le _ }, + rw [a1, one_opow, one_opow] }, + rwa [opow_le_opow_iff_right a1, one_le_iff_pos] +end + +theorem right_le_opow {a : ordinal} (b) (a1 : 1 < a) : b ≤ a ^ b := +(opow_is_normal a1).self_le _ + +theorem opow_lt_opow_left_of_succ {a b c : ordinal} + (ab : a < b) : a ^ succ c < b ^ succ c := +by { rw [opow_succ, opow_succ], exact + (mul_le_mul_right' (opow_le_opow_left c ab.le) a).trans_lt + (mul_lt_mul_of_pos_left ab (opow_pos c ((ordinal.zero_le a).trans_lt ab))) } + +theorem opow_add (a b c : ordinal) : a ^ (b + c) = a ^ b * a ^ c := +begin + rcases eq_or_ne a 0 with rfl | a0, + { rcases eq_or_ne c 0 with rfl | c0, { simp }, + have : b + c ≠ 0 := ((ordinal.pos_iff_ne_zero.2 c0).trans_le (le_add_left _ _)).ne', + simp only [zero_opow c0, zero_opow this, mul_zero] }, + rcases eq_or_lt_of_le (one_le_iff_ne_zero.2 a0) with rfl | a1, + { simp only [one_opow, mul_one] }, + apply limit_rec_on c, + { simp }, + { intros c IH, + rw [add_succ, opow_succ, IH, opow_succ, mul_assoc] }, + { intros c l IH, + refine eq_of_forall_ge_iff (λ d, (((opow_is_normal a1).trans + (add_is_normal b)).limit_le l).trans _), + dsimp only [function.comp], + simp only [IH] {contextual := tt}, + exact (((mul_is_normal $ opow_pos b (ordinal.pos_iff_ne_zero.2 a0)).trans + (opow_is_normal a1)).limit_le l).symm } +end + +theorem opow_one_add (a b : ordinal) : a ^ (1 + b) = a * a ^ b := +by rw [opow_add, opow_one] + +theorem opow_dvd_opow (a) {b c : ordinal} + (h : b ≤ c) : a ^ b ∣ a ^ c := +by { rw [← ordinal.add_sub_cancel_of_le h, opow_add], apply dvd_mul_right } + +theorem opow_dvd_opow_iff {a b c : ordinal} + (a1 : 1 < a) : a ^ b ∣ a ^ c ↔ b ≤ c := +⟨λ h, le_of_not_lt $ λ hn, + not_le_of_lt ((opow_lt_opow_iff_right a1).2 hn) $ + le_of_dvd (opow_ne_zero _ $ one_le_iff_ne_zero.1 $ a1.le) h, +opow_dvd_opow _⟩ + +theorem opow_mul (a b c : ordinal) : a ^ (b * c) = (a ^ b) ^ c := +begin + by_cases b0 : b = 0, {simp only [b0, zero_mul, opow_zero, one_opow]}, + by_cases a0 : a = 0, + { subst a, + by_cases c0 : c = 0, {simp only [c0, mul_zero, opow_zero]}, + simp only [zero_opow b0, zero_opow c0, zero_opow (mul_ne_zero b0 c0)] }, + cases eq_or_lt_of_le (one_le_iff_ne_zero.2 a0) with a1 a1, + { subst a1, simp only [one_opow] }, + apply limit_rec_on c, + { simp only [mul_zero, opow_zero] }, + { intros c IH, + rw [mul_succ, opow_add, IH, opow_succ] }, + { intros c l IH, + refine eq_of_forall_ge_iff (λ d, (((opow_is_normal a1).trans + (mul_is_normal (ordinal.pos_iff_ne_zero.2 b0))).limit_le l).trans _), + dsimp only [function.comp], + simp only [IH] {contextual := tt}, + exact (opow_le_of_limit (opow_ne_zero _ a0) l).symm } +end + +/-! ### Ordinal logarithm -/ + +/-- The ordinal logarithm is the solution `u` to the equation `x = b ^ u * v + w` where `v < b` and + `w < b ^ u`. -/ +@[pp_nodot] def log (b : ordinal) (x : ordinal) : ordinal := +if h : 1 < b then pred (Inf {o | x < b ^ o}) else 0 + +/-- The set in the definition of `log` is nonempty. -/ +theorem log_nonempty {b x : ordinal} (h : 1 < b) : {o | x < b ^ o}.nonempty := +⟨_, succ_le_iff.1 (right_le_opow _ h)⟩ + +theorem log_def {b : ordinal} (h : 1 < b) (x : ordinal) : log b x = pred (Inf {o | x < b ^ o}) := +by simp only [log, dif_pos h] + +theorem log_of_not_one_lt_left {b : ordinal} (h : ¬ 1 < b) (x : ordinal) : log b x = 0 := +by simp only [log, dif_neg h] + +theorem log_of_left_le_one {b : ordinal} (h : b ≤ 1) : ∀ x, log b x = 0 := +log_of_not_one_lt_left h.not_lt + +@[simp] lemma log_zero_left : ∀ b, log 0 b = 0 := +log_of_left_le_one zero_le_one + +@[simp] theorem log_zero_right (b : ordinal) : log b 0 = 0 := +if b1 : 1 < b then begin + rw [log_def b1, ← ordinal.le_zero, pred_le], + apply cInf_le', + dsimp, + rw [succ_zero, opow_one], + exact zero_lt_one.trans b1 +end +else by simp only [log_of_not_one_lt_left b1] + +@[simp] theorem log_one_left : ∀ b, log 1 b = 0 := +log_of_left_le_one le_rfl + +theorem succ_log_def {b x : ordinal} (hb : 1 < b) (hx : x ≠ 0) : + succ (log b x) = Inf {o | x < b ^ o} := +begin + let t := Inf {o | x < b ^ o}, + have : x < b ^ t := Inf_mem (log_nonempty hb), + rcases zero_or_succ_or_limit t with h|h|h, + { refine ((one_le_iff_ne_zero.2 hx).not_lt _).elim, + simpa only [h, opow_zero] }, + { rw [show log b x = pred t, from log_def hb x, + succ_pred_iff_is_succ.2 h] }, + { rcases (lt_opow_of_limit (zero_lt_one.trans hb).ne' h).1 this with ⟨a, h₁, h₂⟩, + exact h₁.not_le.elim ((le_cInf_iff'' (log_nonempty hb)).1 le_rfl a h₂) } +end + +theorem lt_opow_succ_log_self {b : ordinal} (hb : 1 < b) (x : ordinal) : x < b ^ succ (log b x) := +begin + rcases eq_or_ne x 0 with rfl | hx, + { apply opow_pos _ (zero_lt_one.trans hb) }, + { rw succ_log_def hb hx, + exact Inf_mem (log_nonempty hb) } +end + +theorem opow_log_le_self (b) {x : ordinal} (hx : x ≠ 0) : b ^ log b x ≤ x := +begin + rcases eq_or_ne b 0 with rfl | b0, + { rw zero_opow', + refine (sub_le_self _ _).trans (one_le_iff_ne_zero.2 hx) }, + rcases lt_or_eq_of_le (one_le_iff_ne_zero.2 b0) with hb | rfl, + { refine le_of_not_lt (λ h, (lt_succ (log b x)).not_le _), + have := @cInf_le' _ _ {o | x < b ^ o} _ h, + rwa ←succ_log_def hb hx at this }, + { rwa [one_opow, one_le_iff_ne_zero] } +end + +/-- `opow b` and `log b` (almost) form a Galois connection. -/ +theorem opow_le_iff_le_log {b x c : ordinal} (hb : 1 < b) (hx : x ≠ 0) : b ^ c ≤ x ↔ c ≤ log b x := +⟨λ h, le_of_not_lt $ λ hn, + (lt_opow_succ_log_self hb x).not_le $ + ((opow_le_opow_iff_right hb).2 (succ_le_of_lt hn)).trans h, +λ h, ((opow_le_opow_iff_right hb).2 h).trans (opow_log_le_self b hx)⟩ + +theorem lt_opow_iff_log_lt {b x c : ordinal} (hb : 1 < b) (hx : x ≠ 0) : x < b ^ c ↔ log b x < c := +lt_iff_lt_of_le_iff_le (opow_le_iff_le_log hb hx) + +theorem log_pos {b o : ordinal} (hb : 1 < b) (ho : o ≠ 0) (hbo : b ≤ o) : 0 < log b o := +by rwa [←succ_le_iff, succ_zero, ←opow_le_iff_le_log hb ho, opow_one] + +theorem log_eq_zero {b o : ordinal} (hbo : o < b) : log b o = 0 := +begin + rcases eq_or_ne o 0 with rfl | ho, + { exact log_zero_right b }, + cases le_or_lt b 1 with hb hb, + { rcases le_one_iff.1 hb with rfl | rfl, + { exact log_zero_left o }, + { exact log_one_left o } }, + { rwa [←ordinal.le_zero, ←lt_succ_iff, succ_zero, ←lt_opow_iff_log_lt hb ho, opow_one] } +end + +@[mono] theorem log_mono_right (b) {x y : ordinal} (xy : x ≤ y) : log b x ≤ log b y := +if hx : x = 0 then by simp only [hx, log_zero_right, ordinal.zero_le] else +if hb : 1 < b then + (opow_le_iff_le_log hb (lt_of_lt_of_le (ordinal.pos_iff_ne_zero.2 hx) xy).ne').1 $ + (opow_log_le_self _ hx).trans xy +else by simp only [log_of_not_one_lt_left hb, ordinal.zero_le] + +theorem log_le_self (b x : ordinal) : log b x ≤ x := +if hx : x = 0 then by simp only [hx, log_zero_right, ordinal.zero_le] else +if hb : 1 < b then (right_le_opow _ hb).trans (opow_log_le_self b hx) +else by simp only [log_of_not_one_lt_left hb, ordinal.zero_le] + +@[simp] theorem log_one_right (b : ordinal) : log b 1 = 0 := +if hb : 1 < b then log_eq_zero hb else log_of_not_one_lt_left hb 1 + +theorem mod_opow_log_lt_self (b : ordinal) {o : ordinal} (ho : o ≠ 0) : o % b ^ log b o < o := +begin + rcases eq_or_ne b 0 with rfl | hb, + { simpa using ordinal.pos_iff_ne_zero.2 ho }, + { exact (mod_lt _ $ opow_ne_zero _ hb).trans_le (opow_log_le_self _ ho) } +end + +theorem log_mod_opow_log_lt_log_self {b o : ordinal} (hb : 1 < b) (ho : o ≠ 0) (hbo : b ≤ o) : + log b (o % b ^ log b o) < log b o := +begin + cases eq_or_ne (o % b ^ log b o) 0, + { rw [h, log_zero_right], + apply log_pos hb ho hbo }, + { rw [←succ_le_iff, succ_log_def hb h], + apply cInf_le', + apply mod_lt, + rw ←ordinal.pos_iff_ne_zero, + exact opow_pos _ (zero_lt_one.trans hb) } +end + +lemma opow_mul_add_pos {b v : ordinal} (hb : b ≠ 0) (u) (hv : v ≠ 0) (w) : 0 < b ^ u * v + w := +(opow_pos u $ ordinal.pos_iff_ne_zero.2 hb).trans_le $ + (le_mul_left _ $ ordinal.pos_iff_ne_zero.2 hv).trans $ le_add_right _ _ + +lemma opow_mul_add_lt_opow_mul_succ {b u w : ordinal} (v : ordinal) (hw : w < b ^ u) : + b ^ u * v + w < b ^ u * (succ v) := +by rwa [mul_succ, add_lt_add_iff_left] + +lemma opow_mul_add_lt_opow_succ {b u v w : ordinal} (hvb : v < b) (hw : w < b ^ u) : + b ^ u * v + w < b ^ (succ u) := +begin + convert (opow_mul_add_lt_opow_mul_succ v hw).trans_le (mul_le_mul_left' (succ_le_of_lt hvb) _), + exact opow_succ b u +end + +theorem log_opow_mul_add {b u v w : ordinal} (hb : 1 < b) (hv : v ≠ 0) (hvb : v < b) + (hw : w < b ^ u) : log b (b ^ u * v + w) = u := +begin + have hne' := (opow_mul_add_pos (zero_lt_one.trans hb).ne' u hv w).ne', + by_contra' hne, + cases lt_or_gt_of_ne hne with h h, + { rw ←lt_opow_iff_log_lt hb hne' at h, + exact h.not_le ((le_mul_left _ (ordinal.pos_iff_ne_zero.2 hv)).trans (le_add_right _ _)) }, + { change _ < _ at h, + rw [←succ_le_iff, ←opow_le_iff_le_log hb hne'] at h, + exact (not_lt_of_le h) (opow_mul_add_lt_opow_succ hvb hw) } +end + +theorem log_opow {b : ordinal} (hb : 1 < b) (x : ordinal) : log b (b ^ x) = x := +begin + convert log_opow_mul_add hb zero_ne_one.symm hb (opow_pos x (zero_lt_one.trans hb)), + rw [add_zero, mul_one] +end + +theorem div_opow_log_lt {b : ordinal} (o : ordinal) (hb : 1 < b) : o / b ^ log b o < b := +begin + rw [div_lt (opow_pos _ (zero_lt_one.trans hb)).ne', ←opow_succ], + exact lt_opow_succ_log_self hb o +end + +theorem add_log_le_log_mul {x y : ordinal} (b : ordinal) (hx : x ≠ 0) (hy : y ≠ 0) : + log b x + log b y ≤ log b (x * y) := +begin + by_cases hb : 1 < b, + { rw [←opow_le_iff_le_log hb (mul_ne_zero hx hy), opow_add], + exact mul_le_mul' (opow_log_le_self b hx) (opow_log_le_self b hy) }, + simp only [log_of_not_one_lt_left hb, zero_add] +end + +/-! ### Interaction with `nat.cast` -/ + +@[simp, norm_cast] theorem nat_cast_opow (m : ℕ) : ∀ n : ℕ, ((pow m n : ℕ) : ordinal) = m ^ n +| 0 := by simp +| (n+1) := by rw [pow_succ', nat_cast_mul, nat_cast_opow, nat.cast_succ, add_one_eq_succ, opow_succ] + +local infixr (name := ordinal.pow) ^ := @pow ordinal ordinal ordinal.has_pow +theorem sup_opow_nat {o : ordinal} (ho : 0 < o) : sup (λ n : ℕ, o ^ n) = o ^ ω := +begin + rcases lt_or_eq_of_le (one_le_iff_pos.2 ho) with ho₁ | rfl, + { exact (opow_is_normal ho₁).apply_omega }, + { rw one_opow, + refine le_antisymm (sup_le (λ n, by rw one_opow)) _, + convert le_sup _ 0, + rw [nat.cast_zero, opow_zero] } +end + +end ordinal + +namespace tactic +open ordinal positivity + +/-- Extension for the `positivity` tactic: `ordinal.opow` takes positive values on positive inputs. +-/ +@[positivity] +meta def positivity_opow : expr → tactic strictness +| `(@has_pow.pow _ _ %%inst %%a %%b) := do + strictness_a ← core a, + match strictness_a with + | positive p := positive <$> mk_app ``opow_pos [b, p] + | _ := failed -- We already know that `0 ≤ x` for all `x : ordinal` + end +| _ := failed + +end tactic diff --git a/src/set_theory/ordinal/fixed_point.lean b/src/set_theory/ordinal/fixed_point.lean index adfdd0d7ec60e..e59d33952814a 100644 --- a/src/set_theory/ordinal/fixed_point.lean +++ b/src/set_theory/ordinal/fixed_point.lean @@ -5,6 +5,7 @@ Authors: Violeta Hernández Palacios, Mario Carneiro -/ import set_theory.ordinal.arithmetic +import set_theory.ordinal.exponential /-! # Fixed points of normal functions From 32253a1a1071173b33dc7d6a218cf722c6feb514 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 15 Feb 2023 17:45:13 +0000 Subject: [PATCH 0481/1291] split(topology/algebra/infinite_sum): Split into four files (#18414) Over the past five years, `topology.algebra.infinite_sum` has grown pretty big. This PR splits off a third of it to three new files. Precisely, split `topology.algebra.infinite_sum` into: * `topology.algebra.infinite_sum.basic`: Definitions and theory in monoids. All this will be multiplicativised in #18405. * `topology.algebra.infinite_sum.ring`: Interaction of infinite sums and (scalar) multiplication. This mostly cannot be multiplicativised. * `topology.algebra.infinite_sum.order`: Interaction of infinite sums and order. Most of this will be multiplicativised in #18405. Incidentally, this brings down some files' imports by quite a lot. Also move the `star` and `mul_opposite` material to the end of the file to facilitate multiplicativisation in #18405. Johannes wrote some of all these files, except `topology.algebra.infinite_sum.real` whose first lemma I could trace back as coming from #753, with the others coming from #1739. --- src/analysis/analytic/basic.lean | 2 +- src/analysis/analytic/isolated_zeros.lean | 1 - src/analysis/complex/basic.lean | 2 +- .../normed_space/finite_dimension.lean | 2 +- src/measure_theory/measure/outer_measure.lean | 1 - src/number_theory/l_series.lean | 2 +- .../basic.lean} | 724 ++---------------- .../module.lean} | 21 +- src/topology/algebra/infinite_sum/order.lean | 262 +++++++ src/topology/algebra/infinite_sum/real.lean | 90 +++ src/topology/algebra/infinite_sum/ring.lean | 217 ++++++ src/topology/continuous_function/algebra.lean | 14 +- src/topology/instances/ennreal.lean | 2 + src/topology/instances/matrix.lean | 2 +- src/topology/instances/nnreal.lean | 5 +- src/topology/instances/triv_sq_zero_ext.lean | 2 +- 16 files changed, 691 insertions(+), 658 deletions(-) rename src/topology/algebra/{infinite_sum.lean => infinite_sum/basic.lean} (67%) rename src/topology/algebra/{module/infinite_sum.lean => infinite_sum/module.lean} (80%) create mode 100644 src/topology/algebra/infinite_sum/order.lean create mode 100644 src/topology/algebra/infinite_sum/real.lean create mode 100644 src/topology/algebra/infinite_sum/ring.lean diff --git a/src/analysis/analytic/basic.lean b/src/analysis/analytic/basic.lean index 988541d8d5778..c33669722d9d5 100644 --- a/src/analysis/analytic/basic.lean +++ b/src/analysis/analytic/basic.lean @@ -6,7 +6,7 @@ Authors: Sébastien Gouëzel, Yury Kudryashov import analysis.calculus.formal_multilinear_series import analysis.specific_limits.normed import logic.equiv.fin -import topology.algebra.module.infinite_sum +import topology.algebra.infinite_sum.module /-! # Analytic functions diff --git a/src/analysis/analytic/isolated_zeros.lean b/src/analysis/analytic/isolated_zeros.lean index a46bfea457a1d..8d7b7cffa7e1b 100644 --- a/src/analysis/analytic/isolated_zeros.lean +++ b/src/analysis/analytic/isolated_zeros.lean @@ -7,7 +7,6 @@ import analysis.analytic.basic import analysis.calculus.dslope import analysis.calculus.fderiv_analytic import analysis.calculus.formal_multilinear_series -import topology.algebra.infinite_sum import analysis.analytic.uniqueness /-! diff --git a/src/analysis/complex/basic.lean b/src/analysis/complex/basic.lean index 6d82754acffc0..455d27e524793 100644 --- a/src/analysis/complex/basic.lean +++ b/src/analysis/complex/basic.lean @@ -6,7 +6,7 @@ Authors: Sébastien Gouëzel import data.complex.module import data.complex.exponential import data.is_R_or_C.basic -import topology.algebra.module.infinite_sum +import topology.algebra.infinite_sum.module import topology.instances.real_vector_space /-! diff --git a/src/analysis/normed_space/finite_dimension.lean b/src/analysis/normed_space/finite_dimension.lean index 01b7a84f295b7..8ef4d84dab195 100644 --- a/src/analysis/normed_space/finite_dimension.lean +++ b/src/analysis/normed_space/finite_dimension.lean @@ -10,7 +10,7 @@ import analysis.normed_space.operator_norm import analysis.normed_space.riesz_lemma import linear_algebra.matrix.to_lin import topology.algebra.module.finite_dimension -import topology.algebra.module.infinite_sum +import topology.algebra.infinite_sum.module import topology.instances.matrix /-! diff --git a/src/measure_theory/measure/outer_measure.lean b/src/measure_theory/measure/outer_measure.lean index 5cdfc931f5781..28fbdf03a3618 100644 --- a/src/measure_theory/measure/outer_measure.lean +++ b/src/measure_theory/measure/outer_measure.lean @@ -7,7 +7,6 @@ import analysis.specific_limits.basic import measure_theory.pi_system import data.countable.basic import data.fin.vec_notation -import topology.algebra.infinite_sum /-! # Outer Measures diff --git a/src/number_theory/l_series.lean b/src/number_theory/l_series.lean index 006b39f8105e5..0e6db23043981 100644 --- a/src/number_theory/l_series.lean +++ b/src/number_theory/l_series.lean @@ -6,7 +6,7 @@ Authors: Aaron Anderson import analysis.normed_space.finite_dimension import analysis.p_series import number_theory.arithmetic_function -import topology.algebra.infinite_sum +import topology.algebra.infinite_sum.basic /-! # L-series diff --git a/src/topology/algebra/infinite_sum.lean b/src/topology/algebra/infinite_sum/basic.lean similarity index 67% rename from src/topology/algebra/infinite_sum.lean rename to src/topology/algebra/infinite_sum/basic.lean index 04c92573cb782..1ad434de6e1d2 100644 --- a/src/topology/algebra/infinite_sum.lean +++ b/src/topology/algebra/infinite_sum/basic.lean @@ -3,12 +3,10 @@ Copyright (c) 2017 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl -/ -import algebra.big_operators.intervals -import algebra.big_operators.nat_antidiagonal +import data.nat.parity import logic.encodable.lattice -import topology.algebra.mul_action -import topology.algebra.order.monotone_convergence -import topology.instances.real +import topology.algebra.uniform_group +import topology.algebra.star /-! # Infinite sum over a topological monoid @@ -27,8 +25,8 @@ generally, see `has_sum.tendsto_sum_nat`. -/ noncomputable theory -open finset filter function classical -open_locale topology classical big_operators nnreal +open classical filter finset function +open_locale big_operators classical topology variables {α : Type*} {β : Type*} {γ : Type*} {δ : Type*} @@ -254,56 +252,6 @@ lemma function.surjective.summable_iff_of_has_sum_iff {α' : Type*} [add_comm_mo summable f ↔ summable g := hes.exists.trans $ exists_congr $ @he -section mul_opposite -open mul_opposite - -lemma has_sum.op (hf : has_sum f a) : has_sum (λ a, op (f a)) (op a) := -(hf.map (@op_add_equiv α _) continuous_op : _) - -lemma summable.op (hf : summable f) : summable (op ∘ f) := hf.has_sum.op.summable - -lemma has_sum.unop {f : β → αᵐᵒᵖ} {a : αᵐᵒᵖ} (hf : has_sum f a) : - has_sum (λ a, unop (f a)) (unop a) := -(hf.map (@op_add_equiv α _).symm continuous_unop : _) - -lemma summable.unop {f : β → αᵐᵒᵖ} (hf : summable f) : summable (unop ∘ f) := -hf.has_sum.unop.summable - -@[simp] lemma has_sum_op : has_sum (λ a, op (f a)) (op a) ↔ has_sum f a := -⟨has_sum.unop, has_sum.op⟩ - -@[simp] lemma has_sum_unop {f : β → αᵐᵒᵖ} {a : αᵐᵒᵖ} : - has_sum (λ a, unop (f a)) (unop a) ↔ has_sum f a := -⟨has_sum.op, has_sum.unop⟩ - -@[simp] lemma summable_op : summable (λ a, op (f a)) ↔ summable f := -⟨summable.unop, summable.op⟩ - -@[simp] lemma summable_unop {f : β → αᵐᵒᵖ} : summable (λ a, unop (f a)) ↔ summable f := -⟨summable.op, summable.unop⟩ - -end mul_opposite - -section has_continuous_star -variables [star_add_monoid α] [has_continuous_star α] - -lemma has_sum.star (h : has_sum f a) : has_sum (λ b, star (f b)) (star a) := -by simpa only using h.map (star_add_equiv : α ≃+ α) continuous_star - -lemma summable.star (hf : summable f) : summable (λ b, star (f b)) := -hf.has_sum.star.summable - -lemma summable.of_star (hf : summable (λ b, star (f b))) : summable f := -by simpa only [star_star] using hf.star - -@[simp] lemma summable_star_iff : summable (λ b, star (f b)) ↔ summable f := -⟨summable.of_star, summable.star⟩ - -@[simp] lemma summable_star_iff' : summable (star f) ↔ summable f := -summable_star_iff - -end has_continuous_star - variable [has_continuous_add α] lemma has_sum.add (hf : has_sum f a) (hg : has_sum g b) : has_sum (λb, f b + g b) (a + b) := @@ -562,17 +510,6 @@ lemma tsum_eq_tsum_of_ne_zero_bij {g : γ → α} (i : support g → β) ∑' x, f x = ∑' y, g y := tsum_eq_tsum_of_has_sum_iff_has_sum $ λ _, has_sum_iff_has_sum_of_ne_zero_bij i hi hf hfg -lemma tsum_op : ∑' x, mul_opposite.op (f x) = mul_opposite.op (∑' x, f x) := -begin - by_cases h : summable f, - { exact h.has_sum.op.tsum_eq, }, - { have ho := summable_op.not.mpr h, - rw [tsum_eq_zero_of_not_summable h, tsum_eq_zero_of_not_summable ho, mul_opposite.op_zero] }, -end - -lemma tsum_unop {f : β → αᵐᵒᵖ} : ∑' x, mul_opposite.unop (f x) = mul_opposite.unop (∑' x, f x) := -mul_opposite.op_injective tsum_op.symm - /-! ### `tsum` on subsets -/ @[simp] lemma finset.tsum_subtype (s : finset β) (f : β → α) : @@ -656,19 +593,6 @@ end end has_continuous_add -section has_continuous_star -variables [star_add_monoid α] [has_continuous_star α] - -lemma tsum_star : star (∑' b, f b) = ∑' b, star (f b) := -begin - by_cases hf : summable f, - { exact hf.has_sum.star.tsum_eq.symm, }, - { rw [tsum_eq_zero_of_not_summable hf, tsum_eq_zero_of_not_summable (mt summable.of_star hf), - star_zero] }, -end - -end has_continuous_star - open encodable section encodable @@ -764,33 +688,6 @@ lemma tsum_even_add_odd {f : ℕ → α} (he : summable (λ k, f (2 * k))) end tsum -section prod - -variables [add_comm_monoid α] [topological_space α] [add_comm_monoid γ] [topological_space γ] - -lemma has_sum.prod_mk {f : β → α} {g : β → γ} {a : α} {b : γ} - (hf : has_sum f a) (hg : has_sum g b) : - has_sum (λ x, (⟨f x, g x⟩ : α × γ)) ⟨a, b⟩ := -by simp [has_sum, ← prod_mk_sum, filter.tendsto.prod_mk_nhds hf hg] - -end prod - -section pi -variables {ι : Type*} {π : α → Type*} [∀ x, add_comm_monoid (π x)] [∀ x, topological_space (π x)] - -lemma pi.has_sum {f : ι → ∀ x, π x} {g : ∀ x, π x} : - has_sum f g ↔ ∀ x, has_sum (λ i, f i x) (g x) := -by simp only [has_sum, tendsto_pi_nhds, sum_apply] - -lemma pi.summable {f : ι → ∀ x, π x} : summable f ↔ ∀ x, summable (λ i, f i x) := -by simp only [summable, pi.has_sum, skolem] - -lemma tsum_apply [∀ x, t2_space (π x)] {f : ι → ∀ x, π x}{x : α} (hf : summable f) : - (∑' i, f i) x = ∑' i, f i x := -(pi.has_sum.mp hf.has_sum x).tsum_eq.symm - -end pi - section topological_group variables [add_comm_group α] [topological_space α] [topological_add_group α] variables {f g : β → α} {a a₁ a₂ : α} @@ -1052,306 +949,6 @@ end nat end topological_group -section topological_semiring -variables [non_unital_non_assoc_semiring α] [topological_space α] [topological_semiring α] -variables {f g : β → α} {a a₁ a₂ : α} -lemma has_sum.mul_left (a₂) (h : has_sum f a₁) : has_sum (λb, a₂ * f b) (a₂ * a₁) := -by simpa only using h.map (add_monoid_hom.mul_left a₂) (continuous_const.mul continuous_id) - -lemma has_sum.mul_right (a₂) (hf : has_sum f a₁) : has_sum (λb, f b * a₂) (a₁ * a₂) := -by simpa only using hf.map (add_monoid_hom.mul_right a₂) (continuous_id.mul continuous_const) - -lemma summable.mul_left (a) (hf : summable f) : summable (λb, a * f b) := -(hf.has_sum.mul_left _).summable - -lemma summable.mul_right (a) (hf : summable f) : summable (λb, f b * a) := -(hf.has_sum.mul_right _).summable - -section tsum -variables [t2_space α] - -lemma summable.tsum_mul_left (a) (hf : summable f) : ∑'b, a * f b = a * ∑'b, f b := -(hf.has_sum.mul_left _).tsum_eq - -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 - -section const_smul -variables {R : Type*} -[monoid R] -[topological_space α] [add_comm_monoid α] -[distrib_mul_action R α] [has_continuous_const_smul R α] -{f : β → α} - -lemma has_sum.const_smul {a : α} (r : R) (hf : has_sum f a) : has_sum (λ z, r • f z) (r • a) := -hf.map (distrib_mul_action.to_add_monoid_hom α r) (continuous_const_smul r) - -lemma summable.const_smul (r : R) (hf : summable f) : summable (λ z, r • f z) := -(hf.has_sum.const_smul r).summable - -lemma tsum_const_smul [t2_space α] (r : R) (hf : summable f) : ∑' z, r • f z = r • ∑' z, f z := -(hf.has_sum.const_smul r).tsum_eq - -end const_smul - -section smul_const -variables {R : Type*} -[semiring R] [topological_space R] -[topological_space α] [add_comm_monoid α] -[module R α] [has_continuous_smul R α] -{f : β → R} - -lemma has_sum.smul_const {r : R} (hf : has_sum f r) (a : α) : has_sum (λ z, f z • a) (r • a) := -hf.map ((smul_add_hom R α).flip a) (continuous_id.smul continuous_const) - -lemma summable.smul_const (hf : summable f) (a : α) : summable (λ z, f z • a) := -(hf.has_sum.smul_const a).summable - -lemma tsum_smul_const [t2_space α] (hf : summable f) (a : α) : ∑' z, f z • a = (∑' z, f z) • a := -(hf.has_sum.smul_const a).tsum_eq - -end smul_const - -section division_ring - -variables [division_ring α] [topological_space α] [topological_ring α] -{f g : β → α} {a a₁ a₂ : α} - -lemma has_sum.div_const (h : has_sum f a) (b : α) : has_sum (λ x, f x / b) (a / b) := -by simp only [div_eq_mul_inv, h.mul_right b⁻¹] - -lemma summable.div_const (h : summable f) (b : α) : summable (λ x, f x / b) := -(h.has_sum.div_const b).summable - -lemma has_sum_mul_left_iff (h : a₂ ≠ 0) : has_sum (λb, a₂ * f b) (a₂ * a₁) ↔ has_sum f a₁ := -⟨λ H, by simpa only [inv_mul_cancel_left₀ h] using H.mul_left a₂⁻¹, has_sum.mul_left _⟩ - -lemma has_sum_mul_right_iff (h : a₂ ≠ 0) : has_sum (λb, f b * a₂) (a₁ * a₂) ↔ has_sum f a₁ := -⟨λ H, by simpa only [mul_inv_cancel_right₀ h] using H.mul_right a₂⁻¹, has_sum.mul_right _⟩ - -lemma has_sum_div_const_iff (h : a₂ ≠ 0) : has_sum (λb, f b / a₂) (a₁ / a₂) ↔ has_sum f a₁ := -by simpa only [div_eq_mul_inv] using has_sum_mul_right_iff (inv_ne_zero h) - -lemma summable_mul_left_iff (h : a ≠ 0) : summable (λb, a * f b) ↔ summable f := -⟨λ H, by simpa only [inv_mul_cancel_left₀ h] using H.mul_left a⁻¹, λ H, H.mul_left _⟩ - -lemma summable_mul_right_iff (h : a ≠ 0) : summable (λb, f b * a) ↔ summable f := -⟨λ H, by simpa only [mul_inv_cancel_right₀ h] using H.mul_right a⁻¹, λ H, H.mul_right _⟩ - -lemma summable_div_const_iff (h : a ≠ 0) : summable (λb, f b / a) ↔ summable f := -by simpa only [div_eq_mul_inv] using summable_mul_right_iff (inv_ne_zero h) - -lemma tsum_mul_left [t2_space α] : (∑' x, a * f x) = a * ∑' x, f x := -if hf : summable f then hf.tsum_mul_left a -else if ha : a = 0 then by simp [ha] -else by rw [tsum_eq_zero_of_not_summable hf, - tsum_eq_zero_of_not_summable (mt (summable_mul_left_iff ha).mp hf), mul_zero] - -lemma tsum_mul_right [t2_space α] : (∑' x, f x * a) = (∑' x, f x) * a := -if hf : summable f then hf.tsum_mul_right a -else if ha : a = 0 then by simp [ha] -else by rw [tsum_eq_zero_of_not_summable hf, - tsum_eq_zero_of_not_summable (mt (summable_mul_right_iff ha).mp hf), zero_mul] - -lemma tsum_div_const [t2_space α] : (∑' x, f x / a) = (∑' x, f x) / a := -by simpa only [div_eq_mul_inv] using tsum_mul_right - -end division_ring - -section order_topology -variables [ordered_add_comm_monoid α] [topological_space α] [order_closed_topology α] -variables {f g : β → α} {a a₁ a₂ : α} - -lemma has_sum_le (h : ∀b, f b ≤ g b) (hf : has_sum f a₁) (hg : has_sum g a₂) : a₁ ≤ a₂ := -le_of_tendsto_of_tendsto' hf hg $ assume s, sum_le_sum $ assume b _, h b - -@[mono] lemma has_sum_mono (hf : has_sum f a₁) (hg : has_sum g a₂) (h : f ≤ g) : a₁ ≤ a₂ := -has_sum_le h hf hg - -lemma has_sum_le_of_sum_le (hf : has_sum f a) (h : ∀ s : finset β, ∑ b in s, f b ≤ a₂) : - a ≤ a₂ := -le_of_tendsto' hf h - -lemma le_has_sum_of_le_sum (hf : has_sum f a) (h : ∀ s : finset β, a₂ ≤ ∑ b in s, f b) : - a₂ ≤ a := -ge_of_tendsto' hf h - -lemma has_sum_le_inj {g : γ → α} (i : β → γ) (hi : injective i) (hs : ∀c∉set.range i, 0 ≤ g c) - (h : ∀b, f b ≤ g (i b)) (hf : has_sum f a₁) (hg : has_sum g a₂) : a₁ ≤ a₂ := -have has_sum (λc, (partial_inv i c).cases_on' 0 f) a₁, -begin - refine (has_sum_iff_has_sum_of_ne_zero_bij (i ∘ coe) _ _ _).2 hf, - { exact assume c₁ c₂ eq, hi eq }, - { intros c hc, - rw [mem_support] at hc, - cases eq : partial_inv i c with b; rw eq at hc, - { contradiction }, - { rw [partial_inv_of_injective hi] at eq, - exact ⟨⟨b, hc⟩, eq⟩ } }, - { assume c, simp [partial_inv_left hi, option.cases_on'] } -end, -begin - refine has_sum_le (assume c, _) this hg, - by_cases c ∈ set.range i, - { rcases h with ⟨b, rfl⟩, - rw [partial_inv_left hi, option.cases_on'], - exact h _ }, - { have : partial_inv i c = none := dif_neg h, - rw [this, option.cases_on'], - exact hs _ h } -end - -lemma tsum_le_tsum_of_inj {g : γ → α} (i : β → γ) (hi : injective i) (hs : ∀c∉set.range i, 0 ≤ g c) - (h : ∀b, f b ≤ g (i b)) (hf : summable f) (hg : summable g) : tsum f ≤ tsum g := -has_sum_le_inj i hi hs h hf.has_sum hg.has_sum - -lemma sum_le_has_sum (s : finset β) (hs : ∀ b∉s, 0 ≤ f b) (hf : has_sum f a) : - ∑ b in s, f b ≤ a := -ge_of_tendsto hf (eventually_at_top.2 ⟨s, λ t hst, - sum_le_sum_of_subset_of_nonneg hst $ λ b hbt hbs, hs b hbs⟩) - -lemma is_lub_has_sum (h : ∀ b, 0 ≤ f b) (hf : has_sum f a) : - is_lub (set.range (λ s : finset β, ∑ b in s, f b)) a := -is_lub_of_tendsto_at_top (finset.sum_mono_set_of_nonneg h) hf - -lemma le_has_sum (hf : has_sum f a) (b : β) (hb : ∀ b' ≠ b, 0 ≤ f b') : f b ≤ a := -calc f b = ∑ b in {b}, f b : finset.sum_singleton.symm -... ≤ a : sum_le_has_sum _ (by { convert hb, simp }) hf - -lemma sum_le_tsum {f : β → α} (s : finset β) (hs : ∀ b∉s, 0 ≤ f b) (hf : summable f) : - ∑ b in s, f b ≤ ∑' b, f b := -sum_le_has_sum s hs hf.has_sum - -lemma le_tsum (hf : summable f) (b : β) (hb : ∀ b' ≠ b, 0 ≤ f b') : f b ≤ ∑' b, f b := -le_has_sum (summable.has_sum hf) b hb - -lemma tsum_le_tsum (h : ∀b, f b ≤ g b) (hf : summable f) (hg : summable g) : ∑'b, f b ≤ ∑'b, g b := -has_sum_le h hf.has_sum hg.has_sum - -@[mono] lemma tsum_mono (hf : summable f) (hg : summable g) (h : f ≤ g) : - ∑' n, f n ≤ ∑' n, g n := -tsum_le_tsum h hf hg - -lemma tsum_le_of_sum_le (hf : summable f) (h : ∀ s : finset β, ∑ b in s, f b ≤ a₂) : - ∑' b, f b ≤ a₂ := -has_sum_le_of_sum_le hf.has_sum h - -lemma tsum_le_of_sum_le' (ha₂ : 0 ≤ a₂) (h : ∀ s : finset β, ∑ b in s, f b ≤ a₂) : - ∑' b, f b ≤ a₂ := -begin - by_cases hf : summable f, - { exact tsum_le_of_sum_le hf h }, - { rw tsum_eq_zero_of_not_summable hf, - exact ha₂ } -end - -lemma has_sum.nonneg (h : ∀ b, 0 ≤ g b) (ha : has_sum g a) : 0 ≤ a := -has_sum_le h has_sum_zero ha - -lemma has_sum.nonpos (h : ∀ b, g b ≤ 0) (ha : has_sum g a) : a ≤ 0 := -has_sum_le h ha has_sum_zero - -lemma tsum_nonneg (h : ∀ b, 0 ≤ g b) : 0 ≤ ∑'b, g b := -begin - by_cases hg : summable g, - { exact hg.has_sum.nonneg h }, - { simp [tsum_eq_zero_of_not_summable hg] } -end - -lemma tsum_nonpos (h : ∀ b, f b ≤ 0) : ∑'b, f b ≤ 0 := -begin - by_cases hf : summable f, - { exact hf.has_sum.nonpos h }, - { simp [tsum_eq_zero_of_not_summable hf] } -end - -end order_topology - -section ordered_topological_group - -variables [ordered_add_comm_group α] [topological_space α] [topological_add_group α] - [order_closed_topology α] {f g : β → α} {a₁ a₂ : α} - -lemma has_sum_lt {i : β} (h : ∀ (b : β), f b ≤ g b) (hi : f i < g i) - (hf : has_sum f a₁) (hg : has_sum g a₂) : - a₁ < a₂ := -have update f i 0 ≤ update g i 0 := update_le_update_iff.mpr ⟨rfl.le, λ i _, h i⟩, -have 0 - f i + a₁ ≤ 0 - g i + a₂ := has_sum_le this (hf.update i 0) (hg.update i 0), -by simpa only [zero_sub, add_neg_cancel_left] using add_lt_add_of_lt_of_le hi this - -@[mono] lemma has_sum_strict_mono (hf : has_sum f a₁) (hg : has_sum g a₂) (h : f < g) : a₁ < a₂ := -let ⟨hle, i, hi⟩ := pi.lt_def.mp h in has_sum_lt hle hi hf hg - -lemma tsum_lt_tsum {i : β} (h : ∀ (b : β), f b ≤ g b) (hi : f i < g i) - (hf : summable f) (hg : summable g) : - ∑' n, f n < ∑' n, g n := -has_sum_lt h hi hf.has_sum hg.has_sum - -@[mono] lemma tsum_strict_mono (hf : summable f) (hg : summable g) (h : f < g) : - ∑' n, f n < ∑' n, g n := -let ⟨hle, i, hi⟩ := pi.lt_def.mp h in tsum_lt_tsum hle hi hf hg - -lemma tsum_pos (hsum : summable g) (hg : ∀ b, 0 ≤ g b) (i : β) (hi : 0 < g i) : - 0 < ∑' b, g b := -by { rw ← tsum_zero, exact tsum_lt_tsum hg hi summable_zero hsum } - -lemma has_sum_zero_iff_of_nonneg (hf : ∀ i, 0 ≤ f i) : has_sum f 0 ↔ f = 0 := -begin - split, - { intros hf', - ext i, - by_contra hi', - have hi : 0 < f i := lt_of_le_of_ne (hf i) (ne.symm hi'), - simpa using has_sum_lt hf hi has_sum_zero hf' }, - { rintros rfl, - exact has_sum_zero }, -end - -end ordered_topological_group - -section canonically_ordered -variables [canonically_ordered_add_monoid α] [topological_space α] [order_closed_topology α] -variables {f : β → α} {a : α} - -lemma le_has_sum' (hf : has_sum f a) (b : β) : f b ≤ a := -le_has_sum hf b $ λ _ _, zero_le _ - -lemma le_tsum' (hf : summable f) (b : β) : f b ≤ ∑' b, f b := -le_tsum hf b $ λ _ _, zero_le _ - -lemma has_sum_zero_iff : has_sum f 0 ↔ ∀ x, f x = 0 := -begin - refine ⟨_, λ h, _⟩, - { contrapose!, - exact λ ⟨x, hx⟩ h, irrefl _ (lt_of_lt_of_le (pos_iff_ne_zero.2 hx) (le_has_sum' h x)) }, - { convert has_sum_zero, - exact funext h } -end - -lemma tsum_eq_zero_iff (hf : summable f) : ∑' i, f i = 0 ↔ ∀ x, f x = 0 := -by rw [←has_sum_zero_iff, hf.has_sum_iff] - -lemma tsum_ne_zero_iff (hf : summable f) : ∑' i, f i ≠ 0 ↔ ∃ x, f x ≠ 0 := -by rw [ne.def, tsum_eq_zero_iff hf, not_forall] - -lemma is_lub_has_sum' (hf : has_sum f a) : is_lub (set.range (λ s : finset β, ∑ b in s, f b)) a := -is_lub_of_tendsto_at_top (finset.sum_mono_set f) hf - -end canonically_ordered - section uniform_group variables [add_comm_group α] [uniform_space α] @@ -1529,275 +1126,124 @@ end lemma summable.tendsto_at_top_zero {f : ℕ → G} (hf : summable f) : tendsto f at_top (𝓝 0) := by { rw ←nat.cofinite_eq_at_top, exact hf.tendsto_cofinite_zero } -lemma summable.tendsto_top_of_pos {α : Type*} - [linear_ordered_field α] [topological_space α] [order_topology α] {f : ℕ → α} - (hf : summable f⁻¹) (hf' : ∀ n, 0 < f n) : tendsto f at_top at_top := -begin - rw [show f = f⁻¹⁻¹, by { ext, simp }], - apply filter.tendsto.inv_tendsto_zero, - apply tendsto_nhds_within_of_tendsto_nhds_of_eventually_within _ - (summable.tendsto_at_top_zero hf), - rw eventually_iff_exists_mem, - refine ⟨set.Ioi 0, Ioi_mem_at_top _, λ _ _, _⟩, - rw [set.mem_Ioi, inv_eq_one_div, one_div, pi.inv_apply, _root_.inv_pos], - exact hf' _, -end - end topological_group -section preorder - -variables {E : Type*} [preorder E] [add_comm_monoid E] - [topological_space E] [order_closed_topology E] [t2_space E] - -lemma tsum_le_of_sum_range_le {f : ℕ → E} {c : E} (hsum : summable f) - (h : ∀ n, ∑ i in finset.range n, f i ≤ c) : ∑' n, f n ≤ c := -let ⟨l, hl⟩ := hsum in hl.tsum_eq.symm ▸ le_of_tendsto' hl.tendsto_sum_nat h - -end preorder +section const_smul +variables [monoid γ] [topological_space α] [add_comm_monoid α] [distrib_mul_action γ α] + [has_continuous_const_smul γ α] {f : β → α} -section linear_order +lemma has_sum.const_smul {a : α} (b : γ) (hf : has_sum f a) : has_sum (λ i, b • f i) (b • a) := +hf.map (distrib_mul_action.to_add_monoid_hom α _) $ continuous_const_smul _ -/-! For infinite sums taking values in a linearly ordered monoid, the existence of a least upper -bound for the finite sums is a criterion for summability. +lemma summable.const_smul (b : γ) (hf : summable f) : summable (λ i, b • f i) := +(hf.has_sum.const_smul _).summable -This criterion is useful when applied in a linearly ordered monoid which is also a complete or -conditionally complete linear order, such as `ℝ`, `ℝ≥0`, `ℝ≥0∞`, because it is then easy to check -the existence of a least upper bound. --/ +lemma tsum_const_smul [t2_space α] (b : γ) (hf : summable f) : ∑' i, b • f i = b • ∑' i, f i := +(hf.has_sum.const_smul _).tsum_eq -lemma has_sum_of_is_lub_of_nonneg [linear_ordered_add_comm_monoid β] [topological_space β] - [order_topology β] {f : α → β} (b : β) (h : ∀ b, 0 ≤ f b) - (hf : is_lub (set.range (λ s, ∑ a in s, f a)) b) : - has_sum f b := -tendsto_at_top_is_lub (finset.sum_mono_set_of_nonneg h) hf - -lemma has_sum_of_is_lub [canonically_linear_ordered_add_monoid β] [topological_space β] - [order_topology β] {f : α → β} (b : β) (hf : is_lub (set.range (λ s, ∑ a in s, f a)) b) : - has_sum f b := -tendsto_at_top_is_lub (finset.sum_mono_set f) hf - -lemma summable_abs_iff [linear_ordered_add_comm_group β] [uniform_space β] - [uniform_add_group β] [complete_space β] {f : α → β} : - summable (λ x, |f x|) ↔ summable f := -have h1 : ∀ x : {x | 0 ≤ f x}, |f x| = f x := λ x, abs_of_nonneg x.2, -have h2 : ∀ x : {x | 0 ≤ f x}ᶜ, |f x| = -f x := λ x, abs_of_neg (not_le.1 x.2), -calc summable (λ x, |f x|) ↔ - summable (λ x : {x | 0 ≤ f x}, |f x|) ∧ summable (λ x : {x | 0 ≤ f x}ᶜ, |f x|) : - summable_subtype_and_compl.symm -... ↔ summable (λ x : {x | 0 ≤ f x}, f x) ∧ summable (λ x : {x | 0 ≤ f x}ᶜ, -f x) : - by simp only [h1, h2] -... ↔ _ : by simp only [summable_neg_iff, summable_subtype_and_compl] - -alias summable_abs_iff ↔ summable.of_abs summable.abs - -lemma finite_of_summable_const [linear_ordered_add_comm_group β] [archimedean β] - [topological_space β] [order_closed_topology β] {b : β} (hb : 0 < b) - (hf : summable (λ a : α, b)) : - set.finite (set.univ : set α) := -begin - have H : ∀ s : finset α, s.card • b ≤ ∑' a : α, b, - { intros s, - simpa using sum_le_has_sum s (λ a ha, hb.le) hf.has_sum }, - obtain ⟨n, hn⟩ := archimedean.arch (∑' a : α, b) hb, - have : ∀ s : finset α, s.card ≤ n, - { intros s, - simpa [nsmul_le_nsmul_iff hb] using (H s).trans hn }, - haveI : fintype α := fintype_of_finset_card_le n this, - exact set.finite_univ -end +end const_smul -end linear_order +/-! ### Product and pi types -/ -section cauchy_seq -open filter +section prod +variables [add_comm_monoid α] [topological_space α] [add_comm_monoid γ] [topological_space γ] -/-- If the extended distance between consecutive points of a sequence is estimated -by a summable series of `nnreal`s, then the original sequence is a Cauchy sequence. -/ -lemma cauchy_seq_of_edist_le_of_summable [pseudo_emetric_space α] {f : ℕ → α} (d : ℕ → ℝ≥0) - (hf : ∀ n, edist (f n) (f n.succ) ≤ d n) (hd : summable d) : cauchy_seq f := -begin - refine emetric.cauchy_seq_iff_nnreal.2 (λ ε εpos, _), - -- Actually we need partial sums of `d` to be a Cauchy sequence - replace hd : cauchy_seq (λ (n : ℕ), ∑ x in range n, d x) := - let ⟨_, H⟩ := hd in H.tendsto_sum_nat.cauchy_seq, - -- Now we take the same `N` as in one of the definitions of a Cauchy sequence - refine (metric.cauchy_seq_iff'.1 hd ε (nnreal.coe_pos.2 εpos)).imp (λ N hN n hn, _), - have hsum := hN n hn, - -- We simplify the known inequality - rw [dist_nndist, nnreal.nndist_eq, ← sum_range_add_sum_Ico _ hn, add_tsub_cancel_left] at hsum, - norm_cast at hsum, - replace hsum := lt_of_le_of_lt (le_max_left _ _) hsum, - rw edist_comm, - -- Then use `hf` to simplify the goal to the same form - apply lt_of_le_of_lt (edist_le_Ico_sum_of_edist_le hn (λ k _ _, hf k)), - assumption_mod_cast -end +lemma has_sum.prod_mk {f : β → α} {g : β → γ} {a : α} {b : γ} + (hf : has_sum f a) (hg : has_sum g b) : + has_sum (λ x, (⟨f x, g x⟩ : α × γ)) ⟨a, b⟩ := +by simp [has_sum, ← prod_mk_sum, filter.tendsto.prod_mk_nhds hf hg] -/-- If the distance between consecutive points of a sequence is estimated by a summable series, -then the original sequence is a Cauchy sequence. -/ -lemma cauchy_seq_of_dist_le_of_summable [pseudo_metric_space α] {f : ℕ → α} (d : ℕ → ℝ) - (hf : ∀ n, dist (f n) (f n.succ) ≤ d n) (hd : summable d) : cauchy_seq f := -begin - refine metric.cauchy_seq_iff'.2 (λε εpos, _), - replace hd : cauchy_seq (λ (n : ℕ), ∑ x in range n, d x) := - let ⟨_, H⟩ := hd in H.tendsto_sum_nat.cauchy_seq, - refine (metric.cauchy_seq_iff'.1 hd ε εpos).imp (λ N hN n hn, _), - have hsum := hN n hn, - rw [real.dist_eq, ← sum_Ico_eq_sub _ hn] at hsum, - calc dist (f n) (f N) = dist (f N) (f n) : dist_comm _ _ - ... ≤ ∑ x in Ico N n, d x : dist_le_Ico_sum_of_dist_le hn (λ k _ _, hf k) - ... ≤ |∑ x in Ico N n, d x| : le_abs_self _ - ... < ε : hsum -end +end prod -lemma cauchy_seq_of_summable_dist [pseudo_metric_space α] {f : ℕ → α} - (h : summable (λn, dist (f n) (f n.succ))) : cauchy_seq f := -cauchy_seq_of_dist_le_of_summable _ (λ _, le_rfl) h +section pi +variables {ι : Type*} {π : α → Type*} [∀ x, add_comm_monoid (π x)] [∀ x, topological_space (π x)] -lemma dist_le_tsum_of_dist_le_of_tendsto [pseudo_metric_space α] {f : ℕ → α} (d : ℕ → ℝ) - (hf : ∀ n, dist (f n) (f n.succ) ≤ d n) (hd : summable d) {a : α} (ha : tendsto f at_top (𝓝 a)) - (n : ℕ) : - dist (f n) a ≤ ∑' m, d (n + m) := -begin - refine le_of_tendsto (tendsto_const_nhds.dist ha) - (eventually_at_top.2 ⟨n, λ m hnm, _⟩), - refine le_trans (dist_le_Ico_sum_of_dist_le hnm (λ k _ _, hf k)) _, - rw [sum_Ico_eq_sum_range], - refine sum_le_tsum (range _) (λ _ _, le_trans dist_nonneg (hf _)) _, - exact hd.comp_injective (add_right_injective n) -end +lemma pi.has_sum {f : ι → ∀ x, π x} {g : ∀ x, π x} : + has_sum f g ↔ ∀ x, has_sum (λ i, f i x) (g x) := +by simp only [has_sum, tendsto_pi_nhds, sum_apply] -lemma dist_le_tsum_of_dist_le_of_tendsto₀ [pseudo_metric_space α] {f : ℕ → α} (d : ℕ → ℝ) - (hf : ∀ n, dist (f n) (f n.succ) ≤ d n) (hd : summable d) {a : α} (ha : tendsto f at_top (𝓝 a)) : - dist (f 0) a ≤ tsum d := -by simpa only [zero_add] using dist_le_tsum_of_dist_le_of_tendsto d hf hd ha 0 +lemma pi.summable {f : ι → ∀ x, π x} : summable f ↔ ∀ x, summable (λ i, f i x) := +by simp only [summable, pi.has_sum, skolem] -lemma dist_le_tsum_dist_of_tendsto [pseudo_metric_space α] {f : ℕ → α} - (h : summable (λn, dist (f n) (f n.succ))) {a : α} (ha : tendsto f at_top (𝓝 a)) (n) : - dist (f n) a ≤ ∑' m, dist (f (n+m)) (f (n+m).succ) := -show dist (f n) a ≤ ∑' m, (λx, dist (f x) (f x.succ)) (n + m), from -dist_le_tsum_of_dist_le_of_tendsto (λ n, dist (f n) (f n.succ)) (λ _, le_rfl) h ha n +lemma tsum_apply [∀ x, t2_space (π x)] {f : ι → ∀ x, π x}{x : α} (hf : summable f) : + (∑' i, f i) x = ∑' i, f i x := +(pi.has_sum.mp hf.has_sum x).tsum_eq.symm -lemma dist_le_tsum_dist_of_tendsto₀ [pseudo_metric_space α] {f : ℕ → α} - (h : summable (λn, dist (f n) (f n.succ))) {a : α} (ha : tendsto f at_top (𝓝 a)) : - dist (f 0) a ≤ ∑' n, dist (f n) (f n.succ) := -by simpa only [zero_add] using dist_le_tsum_dist_of_tendsto h ha 0 +end pi -end cauchy_seq +/-! ### Multiplicative opposite -/ -/-! ## Multipliying two infinite sums +section mul_opposite +open mul_opposite +variables [add_comm_monoid α] [topological_space α] {f : β → α} {a : α} -In this section, we prove various results about `(∑' x : β, f x) * (∑' y : γ, g y)`. Note that we -always assume that the family `λ x : β × γ, f x.1 * g x.2` is summable, since there is no way to -deduce this from the summmabilities of `f` and `g` in general, but if you are working in a normed -space, you may want to use the analogous lemmas in `analysis/normed_space/basic` -(e.g `tsum_mul_tsum_of_summable_norm`). +lemma has_sum.op (hf : has_sum f a) : has_sum (λ a, op (f a)) (op a) := +(hf.map (@op_add_equiv α _) continuous_op : _) -We first establish results about arbitrary index types, `β` and `γ`, and then we specialize to -`β = γ = ℕ` to prove the Cauchy product formula (see `tsum_mul_tsum_eq_tsum_sum_antidiagonal`). +lemma summable.op (hf : summable f) : summable (op ∘ f) := hf.has_sum.op.summable -### Arbitrary index types --/ +lemma has_sum.unop {f : β → αᵐᵒᵖ} {a : αᵐᵒᵖ} (hf : has_sum f a) : + has_sum (λ a, unop (f a)) (unop a) := +(hf.map (@op_add_equiv α _).symm continuous_unop : _) -section tsum_mul_tsum +lemma summable.unop {f : β → αᵐᵒᵖ} (hf : summable f) : summable (unop ∘ f) := +hf.has_sum.unop.summable -variables [topological_space α] [t3_space α] [non_unital_non_assoc_semiring α] - [topological_semiring α] {f : β → α} {g : γ → α} {s t u : α} +@[simp] lemma has_sum_op : has_sum (λ a, op (f a)) (op a) ↔ has_sum f a := +⟨has_sum.unop, has_sum.op⟩ -lemma has_sum.mul_eq (hf : has_sum f s) (hg : has_sum g t) - (hfg : has_sum (λ (x : β × γ), f x.1 * g x.2) u) : - s * t = u := -have key₁ : has_sum (λ b, f b * t) (s * t), - from hf.mul_right t, -have this : ∀ b : β, has_sum (λ c : γ, f b * g c) (f b * t), - from λ b, hg.mul_left (f b), -have key₂ : has_sum (λ b, f b * t) u, - from has_sum.prod_fiberwise hfg this, -key₁.unique key₂ +@[simp] lemma has_sum_unop {f : β → αᵐᵒᵖ} {a : αᵐᵒᵖ} : + has_sum (λ a, unop (f a)) (unop a) ↔ has_sum f a := +⟨has_sum.op, has_sum.unop⟩ -lemma has_sum.mul (hf : has_sum f s) (hg : has_sum g t) - (hfg : summable (λ (x : β × γ), f x.1 * g x.2)) : - has_sum (λ (x : β × γ), f x.1 * g x.2) (s * t) := -let ⟨u, hu⟩ := hfg in -(hf.mul_eq hg hu).symm ▸ hu +@[simp] lemma summable_op : summable (λ a, op (f a)) ↔ summable f := ⟨summable.unop, summable.op⟩ -/-- Product of two infinites sums indexed by arbitrary types. - See also `tsum_mul_tsum_of_summable_norm` if `f` and `g` are abolutely summable. -/ -lemma tsum_mul_tsum (hf : summable f) (hg : summable g) - (hfg : summable (λ (x : β × γ), f x.1 * g x.2)) : - (∑' x, f x) * (∑' y, g y) = (∑' z : β × γ, f z.1 * g z.2) := -hf.has_sum.mul_eq hg.has_sum hfg.has_sum +@[simp] lemma summable_unop {f : β → αᵐᵒᵖ} : summable (λ a, unop (f a)) ↔ summable f := +⟨summable.op, summable.unop⟩ -end tsum_mul_tsum +variables [t2_space α] -section cauchy_product +lemma tsum_op : ∑' x, mul_opposite.op (f x) = mul_opposite.op (∑' x, f x) := +begin + by_cases h : summable f, + { exact h.has_sum.op.tsum_eq }, + { have ho := summable_op.not.mpr h, + rw [tsum_eq_zero_of_not_summable h, tsum_eq_zero_of_not_summable ho, mul_opposite.op_zero] } +end -/-! ### `ℕ`-indexed families (Cauchy product) +lemma tsum_unop {f : β → αᵐᵒᵖ} : ∑' x, mul_opposite.unop (f x) = mul_opposite.unop (∑' x, f x) := +mul_opposite.op_injective tsum_op.symm -We prove two versions of the Cauchy product formula. The first one is -`tsum_mul_tsum_eq_tsum_sum_range`, where the `n`-th term is a sum over `finset.range (n+1)` -involving `nat` subtraction. -In order to avoid `nat` subtraction, we also provide `tsum_mul_tsum_eq_tsum_sum_antidiagonal`, -where the `n`-th term is a sum over all pairs `(k, l)` such that `k+l=n`, which corresponds to the -`finset` `finset.nat.antidiagonal n` -/ +end mul_opposite -variables {f : ℕ → α} {g : ℕ → α} +/-! ### Interaction with the star -/ -open finset +section has_continuous_star +variables [add_comm_monoid α] [topological_space α] [star_add_monoid α] [has_continuous_star α] + {f : β → α} {a : α} -variables [topological_space α] [non_unital_non_assoc_semiring α] +lemma has_sum.star (h : has_sum f a) : has_sum (λ b, star (f b)) (star a) := +by simpa only using h.map (star_add_equiv : α ≃+ α) continuous_star -/- The family `(k, l) : ℕ × ℕ ↦ f k * g l` is summable if and only if the family -`(n, k, l) : Σ (n : ℕ), nat.antidiagonal n ↦ f k * g l` is summable. -/ -lemma summable_mul_prod_iff_summable_mul_sigma_antidiagonal {f g : ℕ → α} : - summable (λ x : ℕ × ℕ, f x.1 * g x.2) ↔ - summable (λ x : (Σ (n : ℕ), nat.antidiagonal n), f (x.2 : ℕ × ℕ).1 * g (x.2 : ℕ × ℕ).2) := -nat.sigma_antidiagonal_equiv_prod.summable_iff.symm +lemma summable.star (hf : summable f) : summable (λ b, star (f b)) := +hf.has_sum.star.summable -variables [t3_space α] [topological_semiring α] +lemma summable.of_star (hf : summable (λ b, star (f b))) : summable f := +by simpa only [star_star] using hf.star -lemma summable_sum_mul_antidiagonal_of_summable_mul {f g : ℕ → α} - (h : summable (λ x : ℕ × ℕ, f x.1 * g x.2)) : - summable (λ n, ∑ kl in nat.antidiagonal n, f kl.1 * g kl.2) := -begin - rw summable_mul_prod_iff_summable_mul_sigma_antidiagonal at h, - conv {congr, funext, rw [← finset.sum_finset_coe, ← tsum_fintype]}, - exact h.sigma' (λ n, (has_sum_fintype _).summable), -end +@[simp] lemma summable_star_iff : summable (λ b, star (f b)) ↔ summable f := +⟨summable.of_star, summable.star⟩ -/-- The Cauchy product formula for the product of two infinites sums indexed by `ℕ`, - expressed by summing on `finset.nat.antidiagonal`. - See also `tsum_mul_tsum_eq_tsum_sum_antidiagonal_of_summable_norm` - if `f` and `g` are absolutely summable. -/ -lemma tsum_mul_tsum_eq_tsum_sum_antidiagonal (hf : summable f) (hg : summable g) - (hfg : summable (λ (x : ℕ × ℕ), f x.1 * g x.2)) : - (∑' n, f n) * (∑' n, g n) = (∑' n, ∑ kl in nat.antidiagonal n, f kl.1 * g kl.2) := -begin - conv_rhs {congr, funext, rw [← finset.sum_finset_coe, ← tsum_fintype]}, - rw [tsum_mul_tsum hf hg hfg, ← nat.sigma_antidiagonal_equiv_prod.tsum_eq (_ : ℕ × ℕ → α)], - exact tsum_sigma' (λ n, (has_sum_fintype _).summable) - (summable_mul_prod_iff_summable_mul_sigma_antidiagonal.mp hfg) -end +@[simp] lemma summable_star_iff' : summable (star f) ↔ summable f := summable_star_iff -lemma summable_sum_mul_range_of_summable_mul {f g : ℕ → α} - (h : summable (λ x : ℕ × ℕ, f x.1 * g x.2)) : - summable (λ n, ∑ k in range (n+1), f k * g (n - k)) := -begin - simp_rw ← nat.sum_antidiagonal_eq_sum_range_succ (λ k l, f k * g l), - exact summable_sum_mul_antidiagonal_of_summable_mul h -end +variables [t2_space α] -/-- The Cauchy product formula for the product of two infinites sums indexed by `ℕ`, - expressed by summing on `finset.range`. - See also `tsum_mul_tsum_eq_tsum_sum_range_of_summable_norm` - if `f` and `g` are absolutely summable. -/ -lemma tsum_mul_tsum_eq_tsum_sum_range (hf : summable f) (hg : summable g) - (hfg : summable (λ (x : ℕ × ℕ), f x.1 * g x.2)) : - (∑' n, f n) * (∑' n, g n) = (∑' n, ∑ k in range (n+1), f k * g (n - k)) := +lemma tsum_star : star (∑' b, f b) = ∑' b, star (f b) := begin - simp_rw ← nat.sum_antidiagonal_eq_sum_range_succ (λ k l, f k * g l), - exact tsum_mul_tsum_eq_tsum_sum_antidiagonal hf hg hfg + by_cases hf : summable f, + { exact hf.has_sum.star.tsum_eq.symm, }, + { rw [tsum_eq_zero_of_not_summable hf, tsum_eq_zero_of_not_summable (mt summable.of_star hf), + star_zero] }, end -end cauchy_product +end has_continuous_star diff --git a/src/topology/algebra/module/infinite_sum.lean b/src/topology/algebra/infinite_sum/module.lean similarity index 80% rename from src/topology/algebra/module/infinite_sum.lean rename to src/topology/algebra/infinite_sum/module.lean index 4360e47ab5dd1..5ec77641df867 100644 --- a/src/topology/algebra/module/infinite_sum.lean +++ b/src/topology/algebra/infinite_sum/module.lean @@ -3,17 +3,34 @@ Copyright (c) 2020 Heather Macbeth. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Heather Macbeth, Yury Kudryashov, Frédéric Dupuis -/ +import topology.algebra.infinite_sum.basic import topology.algebra.module.basic -import topology.algebra.infinite_sum /-! # Infinite sums in topological vector spaces -/ +variables {ι R R₂ M M₂ : Type*} + +section smul_const +variables [semiring R] [topological_space R] [topological_space M] [add_comm_monoid M] [module R M] + [has_continuous_smul R M] {f : ι → R} + +lemma has_sum.smul_const {r : R} (hf : has_sum f r) (a : M) : has_sum (λ z, f z • a) (r • a) := +hf.map ((smul_add_hom R M).flip a) (continuous_id.smul continuous_const) + +lemma summable.smul_const (hf : summable f) (a : M) : summable (λ z, f z • a) := +(hf.has_sum.smul_const _).summable + +lemma tsum_smul_const [t2_space M] (hf : summable f) (a : M) : ∑' z, f z • a = (∑' z, f z) • a := +(hf.has_sum.smul_const _).tsum_eq + +end smul_const + section has_sum -- Results in this section hold for continuous additive monoid homomorphisms or equivalences but we -- don't have bundled continuous additive homomorphisms. -variables {ι R R₂ M M₂ : Type*} [semiring R] [semiring R₂] [add_comm_monoid M] [module R M] +variables [semiring R] [semiring R₂] [add_comm_monoid M] [module R M] [add_comm_monoid M₂] [module R₂ M₂] [topological_space M] [topological_space M₂] {σ : R →+* R₂} {σ' : R₂ →+* R} [ring_hom_inv_pair σ σ'] [ring_hom_inv_pair σ' σ] diff --git a/src/topology/algebra/infinite_sum/order.lean b/src/topology/algebra/infinite_sum/order.lean new file mode 100644 index 0000000000000..82195ab12f816 --- /dev/null +++ b/src/topology/algebra/infinite_sum/order.lean @@ -0,0 +1,262 @@ +/- +Copyright (c) 2017 Johannes Hölzl. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johannes Hölzl +-/ +import algebra.order.archimedean +import topology.algebra.infinite_sum.basic +import topology.algebra.order.field +import topology.algebra.order.monotone_convergence + +/-! +# Infinite sum in an order + +This file provides lemmas about the interaction of infinite sums and order operations. +-/ + +open finset filter function +open_locale big_operators classical + +variables {ι κ α : Type*} + +section preorder +variables [preorder α] [add_comm_monoid α] [topological_space α] [order_closed_topology α] + [t2_space α] {f : ℕ → α} {c : α} + +lemma tsum_le_of_sum_range_le (hf : summable f) (h : ∀ n, ∑ i in range n, f i ≤ c) : + ∑' n, f n ≤ c := +let ⟨l, hl⟩ := hf in hl.tsum_eq.symm ▸ le_of_tendsto' hl.tendsto_sum_nat h + +end preorder + +section ordered_add_comm_monoid +variables [ordered_add_comm_monoid α] [topological_space α] [order_closed_topology α] {f g : ι → α} + {a a₁ a₂ : α} + +lemma has_sum_le (h : ∀ i, f i ≤ g i) (hf : has_sum f a₁) (hg : has_sum g a₂) : a₁ ≤ a₂ := +le_of_tendsto_of_tendsto' hf hg $ λ s, sum_le_sum $ λ i _, h i + +@[mono] lemma has_sum_mono (hf : has_sum f a₁) (hg : has_sum g a₂) (h : f ≤ g) : a₁ ≤ a₂ := +has_sum_le h hf hg + +lemma has_sum_le_of_sum_le (hf : has_sum f a) (h : ∀ s, ∑ i in s, f i ≤ a₂) : a ≤ a₂ := +le_of_tendsto' hf h + +lemma le_has_sum_of_le_sum (hf : has_sum f a) (h : ∀ s, a₂ ≤ ∑ i in s, f i) : a₂ ≤ a := +ge_of_tendsto' hf h + +lemma has_sum_le_inj {g : κ → α} (e : ι → κ) (he : injective e) (hs : ∀ c ∉ set.range e, 0 ≤ g c) + (h : ∀ i, f i ≤ g (e i)) (hf : has_sum f a₁) (hg : has_sum g a₂) : a₁ ≤ a₂ := +have has_sum (λ c, (partial_inv e c).cases_on' 0 f) a₁, +begin + refine (has_sum_iff_has_sum_of_ne_zero_bij (e ∘ coe) (λ c₁ c₂ hc, he hc) (λ c hc, _) _).2 hf, + { rw mem_support at hc, + cases eq : partial_inv e c with i; rw eq at hc, + { contradiction }, + { rw [partial_inv_of_injective he] at eq, + exact ⟨⟨i, hc⟩, eq⟩ } }, + { rintro c, + simp [partial_inv_left he, option.cases_on'] } +end, +begin + refine has_sum_le (λ c, _) this hg, + obtain ⟨i, rfl⟩ | h := em (c ∈ set.range e), + { rw [partial_inv_left he, option.cases_on'], + exact h _ }, + { have : partial_inv e c = none := dif_neg h, + rw [this, option.cases_on'], + exact hs _ h } +end + +lemma tsum_le_tsum_of_inj {g : κ → α} (e : ι → κ) (he : injective e) + (hs : ∀ c ∉ set.range e, 0 ≤ g c) (h : ∀ i, f i ≤ g (e i)) (hf : summable f) (hg : summable g) : + tsum f ≤ tsum g := +has_sum_le_inj _ he hs h hf.has_sum hg.has_sum + +lemma sum_le_has_sum (s : finset ι) (hs : ∀ i ∉ s, 0 ≤ f i) (hf : has_sum f a) : + ∑ i in s, f i ≤ a := +ge_of_tendsto hf (eventually_at_top.2 ⟨s, λ t hst, + sum_le_sum_of_subset_of_nonneg hst $ λ i hbt hbs, hs i hbs⟩) + +lemma is_lub_has_sum (h : ∀ i, 0 ≤ f i) (hf : has_sum f a) : + is_lub (set.range $ λ s, ∑ i in s, f i) a := +is_lub_of_tendsto_at_top (finset.sum_mono_set_of_nonneg h) hf + +lemma le_has_sum (hf : has_sum f a) (i : ι) (hb : ∀ b' ≠ i, 0 ≤ f b') : f i ≤ a := +calc f i = ∑ i in {i}, f i : finset.sum_singleton.symm +... ≤ a : sum_le_has_sum _ (by { convert hb, simp }) hf + +lemma sum_le_tsum {f : ι → α} (s : finset ι) (hs : ∀ i ∉ s, 0 ≤ f i) (hf : summable f) : + ∑ i in s, f i ≤ ∑' i, f i := +sum_le_has_sum s hs hf.has_sum + +lemma le_tsum (hf : summable f) (i : ι) (hb : ∀ b' ≠ i, 0 ≤ f b') : f i ≤ ∑' i, f i := +le_has_sum (summable.has_sum hf) i hb + +lemma tsum_le_tsum (h : ∀ i, f i ≤ g i) (hf : summable f) (hg : summable g) : + ∑' i, f i ≤ ∑' i, g i := +has_sum_le h hf.has_sum hg.has_sum + +@[mono] lemma tsum_mono (hf : summable f) (hg : summable g) (h : f ≤ g) : + ∑' n, f n ≤ ∑' n, g n := +tsum_le_tsum h hf hg + +lemma tsum_le_of_sum_le (hf : summable f) (h : ∀ s, ∑ i in s, f i ≤ a₂) : ∑' i, f i ≤ a₂ := +has_sum_le_of_sum_le hf.has_sum h + +lemma tsum_le_of_sum_le' (ha₂ : 0 ≤ a₂) (h : ∀ s, ∑ i in s, f i ≤ a₂) : ∑' i, f i ≤ a₂ := +begin + by_cases hf : summable f, + { exact tsum_le_of_sum_le hf h }, + { rw tsum_eq_zero_of_not_summable hf, + exact ha₂ } +end + +lemma has_sum.nonneg (h : ∀ i, 0 ≤ g i) (ha : has_sum g a) : 0 ≤ a := has_sum_le h has_sum_zero ha +lemma has_sum.nonpos (h : ∀ i, g i ≤ 0) (ha : has_sum g a) : a ≤ 0 := has_sum_le h ha has_sum_zero + +lemma tsum_nonneg (h : ∀ i, 0 ≤ g i) : 0 ≤ ∑' i, g i := +begin + by_cases hg : summable g, + { exact hg.has_sum.nonneg h }, + { simp [tsum_eq_zero_of_not_summable hg] } +end + +lemma tsum_nonpos (h : ∀ i, f i ≤ 0) : ∑' i, f i ≤ 0 := +begin + by_cases hf : summable f, + { exact hf.has_sum.nonpos h }, + { simp [tsum_eq_zero_of_not_summable hf] } +end + +end ordered_add_comm_monoid + +section ordered_add_comm_group +variables [ordered_add_comm_group α] [topological_space α] [topological_add_group α] + [order_closed_topology α] {f g : ι → α} {a₁ a₂ : α} {i : ι} + +lemma has_sum_lt (h : f ≤ g) (hi : f i < g i) (hf : has_sum f a₁) (hg : has_sum g a₂) : a₁ < a₂ := +have update f i 0 ≤ update g i 0 := update_le_update_iff.mpr ⟨rfl.le, λ i _, h i⟩, +have 0 - f i + a₁ ≤ 0 - g i + a₂ := has_sum_le this (hf.update i 0) (hg.update i 0), +by simpa only [zero_sub, add_neg_cancel_left] using add_lt_add_of_lt_of_le hi this + +@[mono] lemma has_sum_strict_mono (hf : has_sum f a₁) (hg : has_sum g a₂) (h : f < g) : a₁ < a₂ := +let ⟨hle, i, hi⟩ := pi.lt_def.mp h in has_sum_lt hle hi hf hg + +lemma tsum_lt_tsum (h : f ≤ g) (hi : f i < g i) (hf : summable f) (hg : summable g) : + ∑' n, f n < ∑' n, g n := +has_sum_lt h hi hf.has_sum hg.has_sum + +@[mono] lemma tsum_strict_mono (hf : summable f) (hg : summable g) (h : f < g) : + ∑' n, f n < ∑' n, g n := +let ⟨hle, i, hi⟩ := pi.lt_def.mp h in tsum_lt_tsum hle hi hf hg + +lemma tsum_pos (hsum : summable g) (hg : ∀ i, 0 ≤ g i) (i : ι) (hi : 0 < g i) : 0 < ∑' i, g i := +by { rw ←tsum_zero, exact tsum_lt_tsum hg hi summable_zero hsum } + +lemma has_sum_zero_iff_of_nonneg (hf : ∀ i, 0 ≤ f i) : has_sum f 0 ↔ f = 0 := +begin + refine ⟨λ hf', _, _⟩, + { ext i, + refine (hf i).eq_of_not_gt (λ hi, _), + simpa using has_sum_lt hf hi has_sum_zero hf' }, + { rintro rfl, + exact has_sum_zero } +end + +end ordered_add_comm_group + +section canonically_ordered_add_monoid +variables [canonically_ordered_add_monoid α] [topological_space α] [order_closed_topology α] + {f : ι → α} {a : α} + +lemma le_has_sum' (hf : has_sum f a) (i : ι) : f i ≤ a := le_has_sum hf i $ λ _ _, zero_le _ + +lemma le_tsum' (hf : summable f) (i : ι) : f i ≤ ∑' i, f i := le_tsum hf i $ λ _ _, zero_le _ + +lemma has_sum_zero_iff : has_sum f 0 ↔ ∀ x, f x = 0 := +begin + refine ⟨_, λ h, _⟩, + { contrapose!, + exact λ ⟨x, hx⟩ h, hx (nonpos_iff_eq_zero.1$ le_has_sum' h x) }, + { convert has_sum_zero, + exact funext h } +end + +lemma tsum_eq_zero_iff (hf : summable f) : ∑' i, f i = 0 ↔ ∀ x, f x = 0 := +by rw [←has_sum_zero_iff, hf.has_sum_iff] + +lemma tsum_ne_zero_iff (hf : summable f) : ∑' i, f i ≠ 0 ↔ ∃ x, f x ≠ 0 := +by rw [ne.def, tsum_eq_zero_iff hf, not_forall] + +lemma is_lub_has_sum' (hf : has_sum f a) : is_lub (set.range $ λ s, ∑ i in s, f i) a := +is_lub_of_tendsto_at_top (finset.sum_mono_set f) hf + +end canonically_ordered_add_monoid + +section linear_order + +/-! +For infinite sums taking values in a linearly ordered monoid, the existence of a least upper +bound for the finite sums is a criterion for summability. + +This criterion is useful when applied in a linearly ordered monoid which is also a complete or +conditionally complete linear order, such as `ℝ`, `ℝ≥0`, `ℝ≥0∞`, because it is then easy to check +the existence of a least upper bound. +-/ + +lemma has_sum_of_is_lub_of_nonneg [linear_ordered_add_comm_monoid α] [topological_space α] + [order_topology α] {f : ι → α} (i : α) (h : ∀ i, 0 ≤ f i) + (hf : is_lub (set.range $ λ s, ∑ i in s, f i) i) : + has_sum f i := +tendsto_at_top_is_lub (finset.sum_mono_set_of_nonneg h) hf + +lemma has_sum_of_is_lub [canonically_linear_ordered_add_monoid α] [topological_space α] + [order_topology α] {f : ι → α} (b : α) (hf : is_lub (set.range $ λ s, ∑ i in s, f i) b) : + has_sum f b := +tendsto_at_top_is_lub (finset.sum_mono_set f) hf + +lemma summable_abs_iff [linear_ordered_add_comm_group α] [uniform_space α] [uniform_add_group α] + [complete_space α] {f : ι → α} : + summable (λ x, |f x|) ↔ summable f := +have h1 : ∀ x : {x | 0 ≤ f x}, |f x| = f x := λ x, abs_of_nonneg x.2, +have h2 : ∀ x : {x | 0 ≤ f x}ᶜ, |f x| = -f x := λ x, abs_of_neg (not_le.1 x.2), +calc summable (λ x, |f x|) ↔ + summable (λ x : {x | 0 ≤ f x}, |f x|) ∧ summable (λ x : {x | 0 ≤ f x}ᶜ, |f x|) : + summable_subtype_and_compl.symm +... ↔ summable (λ x : {x | 0 ≤ f x}, f x) ∧ summable (λ x : {x | 0 ≤ f x}ᶜ, -f x) : + by simp only [h1, h2] +... ↔ _ : by simp only [summable_neg_iff, summable_subtype_and_compl] + +alias summable_abs_iff ↔ summable.of_abs summable.abs + +--TODO: Change the conclusion to `finite ι` +lemma finite_of_summable_const [linear_ordered_add_comm_group α] [topological_space α] + [archimedean α] [order_closed_topology α] {b : α} (hb : 0 < b) (hf : summable (λ i : ι, b)) : + (set.univ : set ι).finite := +begin + have H : ∀ s : finset ι, s.card • b ≤ ∑' i : ι, b, + { intros s, + simpa using sum_le_has_sum s (λ a ha, hb.le) hf.has_sum }, + obtain ⟨n, hn⟩ := archimedean.arch (∑' i : ι, b) hb, + have : ∀ s : finset ι, s.card ≤ n, + { intros s, + simpa [nsmul_le_nsmul_iff hb] using (H s).trans hn }, + haveI : fintype ι := fintype_of_finset_card_le n this, + exact set.finite_univ +end + +end linear_order + +lemma summable.tendsto_top_of_pos [linear_ordered_field α] [topological_space α] [order_topology α] + {f : ℕ → α} (hf : summable f⁻¹) (hf' : ∀ n, 0 < f n) : tendsto f at_top at_top := +begin + rw ←inv_inv f, + apply filter.tendsto.inv_tendsto_zero, + apply tendsto_nhds_within_of_tendsto_nhds_of_eventually_within _ + (summable.tendsto_at_top_zero hf), + rw eventually_iff_exists_mem, + refine ⟨set.Ioi 0, Ioi_mem_at_top _, λ _ _, _⟩, + rw [set.mem_Ioi, inv_eq_one_div, one_div, pi.inv_apply, _root_.inv_pos], + exact hf' _, +end diff --git a/src/topology/algebra/infinite_sum/real.lean b/src/topology/algebra/infinite_sum/real.lean new file mode 100644 index 0000000000000..ffed456271185 --- /dev/null +++ b/src/topology/algebra/infinite_sum/real.lean @@ -0,0 +1,90 @@ +/- +Copyright (c) 2019 Sébastien Gouëzel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sébastien Gouëzel, Yury Kudryashov +-/ +import topology.algebra.infinite_sum.order +import topology.instances.real + +/-! +# Infinite sum in the reals + +This file provides lemmas about Cauchy sequences in terms of infinite sums. +-/ + +open filter finset +open_locale big_operators nnreal topology + +variables {α : Type*} + +/-- If the extended distance between consecutive points of a sequence is estimated +by a summable series of `nnreal`s, then the original sequence is a Cauchy sequence. -/ +lemma cauchy_seq_of_edist_le_of_summable [pseudo_emetric_space α] {f : ℕ → α} (d : ℕ → ℝ≥0) + (hf : ∀ n, edist (f n) (f n.succ) ≤ d n) (hd : summable d) : cauchy_seq f := +begin + refine emetric.cauchy_seq_iff_nnreal.2 (λ ε εpos, _), + -- Actually we need partial sums of `d` to be a Cauchy sequence + replace hd : cauchy_seq (λ (n : ℕ), ∑ x in range n, d x) := + let ⟨_, H⟩ := hd in H.tendsto_sum_nat.cauchy_seq, + -- Now we take the same `N` as in one of the definitions of a Cauchy sequence + refine (metric.cauchy_seq_iff'.1 hd ε (nnreal.coe_pos.2 εpos)).imp (λ N hN n hn, _), + have hsum := hN n hn, + -- We simplify the known inequality + rw [dist_nndist, nnreal.nndist_eq, ← sum_range_add_sum_Ico _ hn, add_tsub_cancel_left] at hsum, + norm_cast at hsum, + replace hsum := lt_of_le_of_lt (le_max_left _ _) hsum, + rw edist_comm, + -- Then use `hf` to simplify the goal to the same form + apply lt_of_le_of_lt (edist_le_Ico_sum_of_edist_le hn (λ k _ _, hf k)), + assumption_mod_cast +end + +variables [pseudo_metric_space α] {f : ℕ → α} {a : α} + +/-- If the distance between consecutive points of a sequence is estimated by a summable series, +then the original sequence is a Cauchy sequence. -/ +lemma cauchy_seq_of_dist_le_of_summable (d : ℕ → ℝ) (hf : ∀ n, dist (f n) (f n.succ) ≤ d n) + (hd : summable d) : cauchy_seq f := +begin + refine metric.cauchy_seq_iff'.2 (λε εpos, _), + replace hd : cauchy_seq (λ (n : ℕ), ∑ x in range n, d x) := + let ⟨_, H⟩ := hd in H.tendsto_sum_nat.cauchy_seq, + refine (metric.cauchy_seq_iff'.1 hd ε εpos).imp (λ N hN n hn, _), + have hsum := hN n hn, + rw [real.dist_eq, ← sum_Ico_eq_sub _ hn] at hsum, + calc dist (f n) (f N) = dist (f N) (f n) : dist_comm _ _ + ... ≤ ∑ x in Ico N n, d x : dist_le_Ico_sum_of_dist_le hn (λ k _ _, hf k) + ... ≤ |∑ x in Ico N n, d x| : le_abs_self _ + ... < ε : hsum +end + +lemma cauchy_seq_of_summable_dist (h : summable (λ n, dist (f n) (f n.succ))) : cauchy_seq f := +cauchy_seq_of_dist_le_of_summable _ (λ _, le_rfl) h + +lemma dist_le_tsum_of_dist_le_of_tendsto (d : ℕ → ℝ) (hf : ∀ n, dist (f n) (f n.succ) ≤ d n) + (hd : summable d) {a : α} (ha : tendsto f at_top (𝓝 a)) (n : ℕ) : + dist (f n) a ≤ ∑' m, d (n + m) := +begin + refine le_of_tendsto (tendsto_const_nhds.dist ha) + (eventually_at_top.2 ⟨n, λ m hnm, _⟩), + refine le_trans (dist_le_Ico_sum_of_dist_le hnm (λ k _ _, hf k)) _, + rw [sum_Ico_eq_sum_range], + refine sum_le_tsum (range _) (λ _ _, le_trans dist_nonneg (hf _)) _, + exact hd.comp_injective (add_right_injective n) +end + +lemma dist_le_tsum_of_dist_le_of_tendsto₀ (d : ℕ → ℝ) (hf : ∀ n, dist (f n) (f n.succ) ≤ d n) + (hd : summable d) (ha : tendsto f at_top (𝓝 a)) : + dist (f 0) a ≤ tsum d := +by simpa only [zero_add] using dist_le_tsum_of_dist_le_of_tendsto d hf hd ha 0 + +lemma dist_le_tsum_dist_of_tendsto (h : summable (λ n, dist (f n) (f n.succ))) + (ha : tendsto f at_top (𝓝 a)) (n) : + dist (f n) a ≤ ∑' m, dist (f (n + m)) (f (n + m).succ) := +show dist (f n) a ≤ ∑' m, (λx, dist (f x) (f x.succ)) (n + m), from +dist_le_tsum_of_dist_le_of_tendsto (λ n, dist (f n) (f n.succ)) (λ _, le_rfl) h ha n + +lemma dist_le_tsum_dist_of_tendsto₀ (h : summable (λ n, dist (f n) (f n.succ))) + (ha : tendsto f at_top (𝓝 a)) : + dist (f 0) a ≤ ∑' n, dist (f n) (f n.succ) := +by simpa only [zero_add] using dist_le_tsum_dist_of_tendsto h ha 0 diff --git a/src/topology/algebra/infinite_sum/ring.lean b/src/topology/algebra/infinite_sum/ring.lean new file mode 100644 index 0000000000000..f3bcb71490d42 --- /dev/null +++ b/src/topology/algebra/infinite_sum/ring.lean @@ -0,0 +1,217 @@ +/- +Copyright (c) 2017 Johannes Hölzl. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johannes Hölzl +-/ +import topology.algebra.infinite_sum.basic +import topology.algebra.ring + +/-! +# Infinite sum in a ring + +This file provides lemmas about the interaction between infinite sums and multiplication. + +## Main results + +* `tsum_mul_tsum_eq_tsum_sum_antidiagonal`: Cauchy product formula +-/ + +open filter finset function +open_locale big_operators classical + +variables {ι κ R α : Type*} + +section non_unital_non_assoc_semiring +variables [non_unital_non_assoc_semiring α] [topological_space α] [topological_semiring α] + {f g : ι → α} {a a₁ a₂ : α} + +lemma has_sum.mul_left (a₂) (h : has_sum f a₁) : has_sum (λ i, a₂ * f i) (a₂ * a₁) := +by simpa only using h.map (add_monoid_hom.mul_left a₂) (continuous_const.mul continuous_id) + +lemma has_sum.mul_right (a₂) (hf : has_sum f a₁) : has_sum (λ i, f i * a₂) (a₁ * a₂) := +by simpa only using hf.map (add_monoid_hom.mul_right a₂) (continuous_id.mul continuous_const) + +lemma summable.mul_left (a) (hf : summable f) : summable (λ i, a * f i) := +(hf.has_sum.mul_left _).summable + +lemma summable.mul_right (a) (hf : summable f) : summable (λ i, f i * a) := +(hf.has_sum.mul_right _).summable + +section tsum +variables [t2_space α] + +lemma summable.tsum_mul_left (a) (hf : summable f) : ∑' i, a * f i = a * ∑' i, f i := +(hf.has_sum.mul_left _).tsum_eq + +lemma summable.tsum_mul_right (a) (hf : summable f) : ∑' i, f i * a = (∑' i, f i) * a := +(hf.has_sum.mul_right _).tsum_eq + +lemma commute.tsum_right (a) (h : ∀ i, commute a (f i)) : commute a (∑' i, f i) := +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 : ∀ i, commute (f i) a) : commute (∑' i, f i) a := +(commute.tsum_right _ $ λ i, (h i).symm).symm + +end tsum +end non_unital_non_assoc_semiring + +section division_semiring +variables [division_semiring α] [topological_space α] [topological_semiring α] {f g : ι → α} + {a a₁ a₂ : α} + +lemma has_sum.div_const (h : has_sum f a) (b : α) : has_sum (λ i, f i / b) (a / b) := +by simp only [div_eq_mul_inv, h.mul_right b⁻¹] + +lemma summable.div_const (h : summable f) (b : α) : summable (λ i, f i / b) := +(h.has_sum.div_const _).summable + +lemma has_sum_mul_left_iff (h : a₂ ≠ 0) : has_sum (λ i, a₂ * f i) (a₂ * a₁) ↔ has_sum f a₁ := +⟨λ H, by simpa only [inv_mul_cancel_left₀ h] using H.mul_left a₂⁻¹, has_sum.mul_left _⟩ + +lemma has_sum_mul_right_iff (h : a₂ ≠ 0) : has_sum (λ i, f i * a₂) (a₁ * a₂) ↔ has_sum f a₁ := +⟨λ H, by simpa only [mul_inv_cancel_right₀ h] using H.mul_right a₂⁻¹, has_sum.mul_right _⟩ + +lemma has_sum_div_const_iff (h : a₂ ≠ 0) : has_sum (λ i, f i / a₂) (a₁ / a₂) ↔ has_sum f a₁ := +by simpa only [div_eq_mul_inv] using has_sum_mul_right_iff (inv_ne_zero h) + +lemma summable_mul_left_iff (h : a ≠ 0) : summable (λ i, a * f i) ↔ summable f := +⟨λ H, by simpa only [inv_mul_cancel_left₀ h] using H.mul_left a⁻¹, λ H, H.mul_left _⟩ + +lemma summable_mul_right_iff (h : a ≠ 0) : summable (λ i, f i * a) ↔ summable f := +⟨λ H, by simpa only [mul_inv_cancel_right₀ h] using H.mul_right a⁻¹, λ H, H.mul_right _⟩ + +lemma summable_div_const_iff (h : a ≠ 0) : summable (λ i, f i / a) ↔ summable f := +by simpa only [div_eq_mul_inv] using summable_mul_right_iff (inv_ne_zero h) + +lemma tsum_mul_left [t2_space α] : (∑' x, a * f x) = a * ∑' x, f x := +if hf : summable f then hf.tsum_mul_left a +else if ha : a = 0 then by simp [ha] +else by rw [tsum_eq_zero_of_not_summable hf, + tsum_eq_zero_of_not_summable (mt (summable_mul_left_iff ha).mp hf), mul_zero] + +lemma tsum_mul_right [t2_space α] : (∑' x, f x * a) = (∑' x, f x) * a := +if hf : summable f then hf.tsum_mul_right a +else if ha : a = 0 then by simp [ha] +else by rw [tsum_eq_zero_of_not_summable hf, + tsum_eq_zero_of_not_summable (mt (summable_mul_right_iff ha).mp hf), zero_mul] + +lemma tsum_div_const [t2_space α] : (∑' x, f x / a) = (∑' x, f x) / a := +by simpa only [div_eq_mul_inv] using tsum_mul_right + +end division_semiring + +/-! +### Multipliying two infinite sums + +In this section, we prove various results about `(∑' x : ι, f x) * (∑' y : κ, g y)`. Note that we +always assume that the family `λ x : ι × κ, f x.1 * g x.2` is summable, since there is no way to +deduce this from the summmabilities of `f` and `g` in general, but if you are working in a normed +space, you may want to use the analogous lemmas in `analysis/normed_space/basic` +(e.g `tsum_mul_tsum_of_summable_norm`). + +We first establish results about arbitrary index types, `ι` and `κ`, and then we specialize to +`ι = κ = ℕ` to prove the Cauchy product formula (see `tsum_mul_tsum_eq_tsum_sum_antidiagonal`). + +#### Arbitrary index types +-/ + +section tsum_mul_tsum +variables [topological_space α] [t3_space α] [non_unital_non_assoc_semiring α] + [topological_semiring α] {f : ι → α} {g : κ → α} {s t u : α} + +lemma has_sum.mul_eq (hf : has_sum f s) (hg : has_sum g t) + (hfg : has_sum (λ (x : ι × κ), f x.1 * g x.2) u) : + s * t = u := +have key₁ : has_sum (λ i, f i * t) (s * t), + from hf.mul_right t, +have this : ∀ i : ι, has_sum (λ c : κ, f i * g c) (f i * t), + from λ i, hg.mul_left (f i), +have key₂ : has_sum (λ i, f i * t) u, + from has_sum.prod_fiberwise hfg this, +key₁.unique key₂ + +lemma has_sum.mul (hf : has_sum f s) (hg : has_sum g t) + (hfg : summable (λ (x : ι × κ), f x.1 * g x.2)) : + has_sum (λ (x : ι × κ), f x.1 * g x.2) (s * t) := +let ⟨u, hu⟩ := hfg in +(hf.mul_eq hg hu).symm ▸ hu + +/-- Product of two infinites sums indexed by arbitrary types. + See also `tsum_mul_tsum_of_summable_norm` if `f` and `g` are abolutely summable. -/ +lemma tsum_mul_tsum (hf : summable f) (hg : summable g) + (hfg : summable (λ (x : ι × κ), f x.1 * g x.2)) : + (∑' x, f x) * (∑' y, g y) = (∑' z : ι × κ, f z.1 * g z.2) := +hf.has_sum.mul_eq hg.has_sum hfg.has_sum + +end tsum_mul_tsum + +/-! +#### `ℕ`-indexed families (Cauchy product) + +We prove two versions of the Cauchy product formula. The first one is +`tsum_mul_tsum_eq_tsum_sum_range`, where the `n`-th term is a sum over `finset.range (n+1)` +involving `nat` subtraction. +In order to avoid `nat` subtraction, we also provide `tsum_mul_tsum_eq_tsum_sum_antidiagonal`, +where the `n`-th term is a sum over all pairs `(k, l)` such that `k+l=n`, which corresponds to the +`finset` `finset.nat.antidiagonal n` +-/ + +section cauchy_product +variables [topological_space α] [non_unital_non_assoc_semiring α] {f g : ℕ → α} + +/- The family `(k, l) : ℕ × ℕ ↦ f k * g l` is summable if and only if the family +`(n, k, l) : Σ (n : ℕ), nat.antidiagonal n ↦ f k * g l` is summable. -/ +lemma summable_mul_prod_iff_summable_mul_sigma_antidiagonal : + summable (λ x : ℕ × ℕ, f x.1 * g x.2) ↔ + summable (λ x : (Σ (n : ℕ), nat.antidiagonal n), f (x.2 : ℕ × ℕ).1 * g (x.2 : ℕ × ℕ).2) := +nat.sigma_antidiagonal_equiv_prod.summable_iff.symm + +variables [t3_space α] [topological_semiring α] + +lemma summable_sum_mul_antidiagonal_of_summable_mul (h : summable (λ x : ℕ × ℕ, f x.1 * g x.2)) : + summable (λ n, ∑ kl in nat.antidiagonal n, f kl.1 * g kl.2) := +begin + rw summable_mul_prod_iff_summable_mul_sigma_antidiagonal at h, + conv {congr, funext, rw [← finset.sum_finset_coe, ← tsum_fintype]}, + exact h.sigma' (λ n, (has_sum_fintype _).summable), +end + +/-- The **Cauchy product formula** for the product of two infinites sums indexed by `ℕ`, expressed +by summing on `finset.nat.antidiagonal`. + +See also `tsum_mul_tsum_eq_tsum_sum_antidiagonal_of_summable_norm` if `f` and `g` are absolutely +summable. -/ +lemma tsum_mul_tsum_eq_tsum_sum_antidiagonal (hf : summable f) (hg : summable g) + (hfg : summable (λ (x : ℕ × ℕ), f x.1 * g x.2)) : + (∑' n, f n) * (∑' n, g n) = (∑' n, ∑ kl in nat.antidiagonal n, f kl.1 * g kl.2) := +begin + conv_rhs {congr, funext, rw [← finset.sum_finset_coe, ← tsum_fintype]}, + rw [tsum_mul_tsum hf hg hfg, ← nat.sigma_antidiagonal_equiv_prod.tsum_eq (_ : ℕ × ℕ → α)], + exact tsum_sigma' (λ n, (has_sum_fintype _).summable) + (summable_mul_prod_iff_summable_mul_sigma_antidiagonal.mp hfg) +end + +lemma summable_sum_mul_range_of_summable_mul (h : summable (λ x : ℕ × ℕ, f x.1 * g x.2)) : + summable (λ n, ∑ k in range (n+1), f k * g (n - k)) := +begin + simp_rw ← nat.sum_antidiagonal_eq_sum_range_succ (λ k l, f k * g l), + exact summable_sum_mul_antidiagonal_of_summable_mul h +end + +/-- The **Cauchy product formula** for the product of two infinites sums indexed by `ℕ`, expressed +by summing on `finset.range`. + +See also `tsum_mul_tsum_eq_tsum_sum_range_of_summable_norm` if `f` and `g` are absolutely summable. +-/ +lemma tsum_mul_tsum_eq_tsum_sum_range (hf : summable f) (hg : summable g) + (hfg : summable (λ (x : ℕ × ℕ), f x.1 * g x.2)) : + (∑' n, f n) * (∑' n, g n) = ∑' n, ∑ k in range (n + 1), f k * g (n - k) := +begin + simp_rw ← nat.sum_antidiagonal_eq_sum_range_succ (λ k l, f k * g l), + exact tsum_mul_tsum_eq_tsum_sum_antidiagonal hf hg hfg +end + +end cauchy_product diff --git a/src/topology/continuous_function/algebra.lean b/src/topology/continuous_function/algebra.lean index 224f10c45667f..0a1bf7d10a6b7 100644 --- a/src/topology/continuous_function/algebra.lean +++ b/src/topology/continuous_function/algebra.lean @@ -3,16 +3,16 @@ Copyright (c) 2019 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison, Nicolò Cavalleri -/ -import topology.algebra.module.basic -import topology.continuous_function.ordered -import topology.algebra.uniform_group -import topology.uniform_space.compact_convergence -import topology.algebra.star -import topology.algebra.infinite_sum import algebra.algebra.pi import algebra.algebra.subalgebra.basic -import tactic.field_simp import algebra.star.star_alg_hom +import tactic.field_simp +import topology.algebra.module.basic +import topology.algebra.infinite_sum.basic +import topology.algebra.star +import topology.algebra.uniform_group +import topology.continuous_function.ordered +import topology.uniform_space.compact_convergence /-! # Algebraic structures over continuous functions diff --git a/src/topology/instances/ennreal.lean b/src/topology/instances/ennreal.lean index 35b53f56688e4..981d10c97a9dc 100644 --- a/src/topology/instances/ennreal.lean +++ b/src/topology/instances/ennreal.lean @@ -6,6 +6,8 @@ Authors: Johannes Hölzl import topology.instances.nnreal import topology.algebra.order.monotone_continuity import analysis.normed.group.basic +import topology.algebra.infinite_sum.real + /-! # Extended non-negative reals -/ diff --git a/src/topology/instances/matrix.lean b/src/topology/instances/matrix.lean index 9a783e84d865b..79c530090c26f 100644 --- a/src/topology/instances/matrix.lean +++ b/src/topology/instances/matrix.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Oliver Nash. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Oliver Nash, Eric Wieser -/ -import topology.algebra.infinite_sum +import topology.algebra.infinite_sum.basic import topology.algebra.ring import topology.algebra.star import linear_algebra.matrix.nonsingular_inverse diff --git a/src/topology/instances/nnreal.lean b/src/topology/instances/nnreal.lean index 2837afba8385a..e14feefae3884 100644 --- a/src/topology/instances/nnreal.lean +++ b/src/topology/instances/nnreal.lean @@ -3,8 +3,9 @@ Copyright (c) 2018 Johan Commelin. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin -/ -import topology.algebra.infinite_sum -import topology.algebra.group_with_zero +import topology.algebra.infinite_sum.order +import topology.algebra.infinite_sum.ring +import topology.instances.real /-! # Topology on `ℝ≥0` diff --git a/src/topology/instances/triv_sq_zero_ext.lean b/src/topology/instances/triv_sq_zero_ext.lean index 2ea5c325c5b64..9f4c07efe72c8 100644 --- a/src/topology/instances/triv_sq_zero_ext.lean +++ b/src/topology/instances/triv_sq_zero_ext.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Eric Wieser -/ import algebra.triv_sq_zero_ext -import topology.algebra.infinite_sum +import topology.algebra.infinite_sum.basic import topology.algebra.module.basic /-! From 5a82b0671532663333e205f422124a98bdfe673f Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 15 Feb 2023 22:40:51 +0000 Subject: [PATCH 0482/1291] =?UTF-8?q?feat(algebra/order/field/basic):=20`a?= =?UTF-8?q?=20=E2=89=A4=20b=20/=20c=20=E2=86=92=20a=20*=20c=20=E2=89=A4=20?= =?UTF-8?q?b`=20(#18367)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit A missing lemma, analogous to the existing `div_le_of_nonneg_of_le_mul` --- src/algebra/order/field/basic.lean | 8 ++++++++ 1 file changed, 8 insertions(+) diff --git a/src/algebra/order/field/basic.lean b/src/algebra/order/field/basic.lean index 126e30e1cd340..d9f98d5104da9 100644 --- a/src/algebra/order/field/basic.lean +++ b/src/algebra/order/field/basic.lean @@ -175,6 +175,14 @@ by { rw [inv_eq_one_div], exact div_lt_iff' ha } lemma div_le_of_nonneg_of_le_mul (hb : 0 ≤ b) (hc : 0 ≤ c) (h : a ≤ c * b) : a / b ≤ c := by { rcases eq_or_lt_of_le hb with rfl|hb', simp [hc], rwa [div_le_iff hb'] } +/-- One direction of `div_le_iff` where `c` is allowed to be `0` (but `b` must be nonnegative) -/ +lemma mul_le_of_nonneg_of_le_div (hb : 0 ≤ b) (hc : 0 ≤ c) (h : a ≤ b / c) : a * c ≤ b := +begin + obtain rfl | hc := hc.eq_or_lt, + { simpa using hb }, + { rwa le_div_iff hc at h } +end + lemma div_le_one_of_le (h : a ≤ b) (hb : 0 ≤ b) : a / b ≤ 1 := div_le_of_nonneg_of_le_mul hb zero_le_one $ by rwa one_mul From 13cd3e89b30352d5b1b7349f5537ea18ba878e40 Mon Sep 17 00:00:00 2001 From: Kyle Miller Date: Thu, 16 Feb 2023 14:01:05 +0000 Subject: [PATCH 0483/1291] feat(combinatorics/simple_graph/connectivity): `concat` and `concat_rec` (#17209) Adds the alternative constructor for walks along with its recursion principle. (Borrowed from lean2/hott.) Co-authored-by: Floris van Doorn --- .../simple_graph/connectivity.lean | 122 ++++++++++++++++++ src/logic/basic.lean | 14 +- 2 files changed, 133 insertions(+), 3 deletions(-) diff --git a/src/combinatorics/simple_graph/connectivity.lean b/src/combinatorics/simple_graph/connectivity.lean index e0b168cc92085..78bb92bca5a91 100644 --- a/src/combinatorics/simple_graph/connectivity.lean +++ b/src/combinatorics/simple_graph/connectivity.lean @@ -151,6 +151,13 @@ def append : Π {u v w : V}, G.walk u v → G.walk v w → G.walk u w | _ _ _ nil q := q | _ _ _ (cons h p) q := cons h (p.append q) +/-- The reversed version of `simple_graph.walk.cons`, concatenating an edge to +the end of a walk. -/ +def concat {u v w : V} (p : G.walk u v) (h : G.adj v w) : G.walk u w := p.append (cons h nil) + +lemma concat_eq_append {u v w : V} (p : G.walk u v) (h : G.adj v w) : + p.concat h = p.append (cons h nil) := rfl + /-- The concatenation of the reverse of the first walk with the second walk. -/ protected def reverse_aux : Π {u v w : V}, G.walk u v → G.walk u w → G.walk v w | _ _ _ nil q := q @@ -215,6 +222,35 @@ lemma append_assoc : Π {u v w x : V} (p : G.walk u v) (q : G.walk v w) (r : G.w (hu : u = u') (hv : v = v') (hw : w = w') : (p.copy hu hv).append (q.copy hv hw) = (p.append q).copy hu hw := by { subst_vars, refl } +lemma concat_nil {u v : V} (h : G.adj u v) : nil.concat h = cons h nil := rfl + +@[simp] lemma concat_cons {u v w x : V} (h : G.adj u v) (p : G.walk v w) (h' : G.adj w x) : + (cons h p).concat h' = cons h (p.concat h') := rfl + +lemma append_concat {u v w x : V} (p : G.walk u v) (q : G.walk v w) (h : G.adj w x) : + p.append (q.concat h) = (p.append q).concat h := append_assoc _ _ _ + +lemma concat_append {u v w x : V} (p : G.walk u v) (h : G.adj v w) (q : G.walk w x) : + (p.concat h).append q = p.append (cons h q) := +by rw [concat_eq_append, ← append_assoc, cons_nil_append] + +/-- A non-trivial `cons` walk is representable as a `concat` walk. -/ +lemma exists_cons_eq_concat : Π {u v w : V} (h : G.adj u v) (p : G.walk v w), + ∃ (x : V) (q : G.walk u x) (h' : G.adj x w), cons h p = q.concat h' +| _ _ _ h nil := ⟨_, nil, h, rfl⟩ +| _ _ _ h (cons h' p) := + begin + obtain ⟨y, q, h'', hc⟩ := exists_cons_eq_concat h' p, + refine ⟨y, cons h q, h'', _⟩, + rw [concat_cons, hc], + end + +/-- A non-trivial `concat` walk is representable as a `cons` walk. -/ +lemma exists_concat_eq_cons : Π {u v w : V} (p : G.walk u v) (h : G.adj v w), + ∃ (x : V) (h' : G.adj u x) (q : G.walk x w), p.concat h = cons h' q +| _ _ _ nil h := ⟨_, h, nil, rfl⟩ +| _ _ _ (cons h' p) h := ⟨_, h', walk.concat p h, concat_cons _ _ _⟩ + @[simp] lemma reverse_nil {u : V} : (nil : G.walk u u).reverse = nil := rfl lemma reverse_singleton {u v : V} (h : G.adj u v) : @@ -250,6 +286,10 @@ by simp [reverse] (p.append q).reverse = q.reverse.append p.reverse := by simp [reverse] +@[simp] lemma reverse_concat {u v w : V} (p : G.walk u v) (h : G.adj v w) : + (p.concat h).reverse = cons (G.symm h) p.reverse := +by simp [concat_eq_append] + @[simp] lemma reverse_reverse : Π {u v : V} (p : G.walk u v), p.reverse.reverse = p | _ _ nil := rfl | _ _ (cons h p) := by simp [reverse_reverse] @@ -268,6 +308,9 @@ by { subst_vars, refl } | _ _ _ nil _ := by simp | _ _ _ (cons _ _) _ := by simp [length_append, add_left_comm, add_comm] +@[simp] lemma length_concat {u v w : V} (p : G.walk u v) (h : G.adj v w) : + (p.concat h).length = p.length + 1 := length_append _ _ + @[simp] protected lemma length_reverse_aux : Π {u v w : V} (p : G.walk u v) (q : G.walk u w), (p.reverse_aux q).length = p.length + q.length | _ _ _ nil _ := by simp! @@ -291,6 +334,76 @@ end @[simp] lemma length_eq_zero_iff {u : V} {p : G.walk u u} : p.length = 0 ↔ p = nil := by cases p; simp +section concat_rec + +variables + {motive : Π (u v : V), G.walk u v → Sort*} + (Hnil : Π {u : V}, motive u u nil) + (Hconcat : Π {u v w : V} (p : G.walk u v) (h : G.adj v w), motive u v p → motive u w (p.concat h)) + +/-- Auxiliary definition for `simple_graph.walk.concat_rec` -/ +def concat_rec_aux : Π {u v : V} (p : G.walk u v), motive v u p.reverse +| _ _ nil := Hnil +| _ _ (cons h p) := eq.rec (Hconcat p.reverse (G.symm h) (concat_rec_aux p)) (reverse_cons h p).symm + +/-- Recursor on walks by inducting on `simple_graph.walk.concat`. + +This is inducting from the opposite end of the walk compared +to `simple_graph.walk.rec`, which inducts on `simple_graph.walk.cons`. -/ +@[elab_as_eliminator] +def concat_rec {u v : V} (p : G.walk u v) : motive u v p := +eq.rec (concat_rec_aux @Hnil @Hconcat p.reverse) (reverse_reverse p) + +@[simp] lemma concat_rec_nil (u : V) : + @concat_rec _ _ motive @Hnil @Hconcat _ _ (nil : G.walk u u) = Hnil := rfl + +@[simp] lemma concat_rec_concat {u v w : V} (p : G.walk u v) (h : G.adj v w) : + @concat_rec _ _ motive @Hnil @Hconcat _ _ (p.concat h) + = Hconcat p h (concat_rec @Hnil @Hconcat p) := +begin + simp only [concat_rec], + apply eq_of_heq, + apply rec_heq_of_heq, + transitivity concat_rec_aux @Hnil @Hconcat (cons h.symm p.reverse), + { congr, simp }, + { rw [concat_rec_aux, rec_heq_iff_heq], + congr; simp [heq_rec_iff_heq], } +end + +end concat_rec + +lemma concat_ne_nil {u v : V} (p : G.walk u v) (h : G.adj v u) : + p.concat h ≠ nil := +by cases p; simp [concat] + +lemma concat_inj {u v v' w : V} + {p : G.walk u v} {h : G.adj v w} {p' : G.walk u v'} {h' : G.adj v' w} + (he : p.concat h = p'.concat h') : + ∃ (hv : v = v'), p.copy rfl hv = p' := +begin + induction p, + { cases p', + { exact ⟨rfl, rfl⟩ }, + { exfalso, + simp only [concat_nil, concat_cons] at he, + obtain ⟨rfl, he⟩ := he, + simp only [heq_iff_eq] at he, + exact concat_ne_nil _ _ he.symm, } }, + { rw concat_cons at he, + cases p', + { exfalso, + simp only [concat_nil] at he, + obtain ⟨rfl, he⟩ := he, + rw [heq_iff_eq] at he, + exact concat_ne_nil _ _ he, }, + { rw concat_cons at he, + simp only at he, + obtain ⟨rfl, he⟩ := he, + rw [heq_iff_eq] at he, + obtain ⟨rfl, rfl⟩ := p_ih he, + exact ⟨rfl, rfl⟩, } } +end + /-- The `support` of a walk is the list of vertices it visits in order. -/ def support : Π {u v : V}, G.walk u v → list V | u v nil := [u] @@ -310,6 +423,9 @@ def edges {u v : V} (p : G.walk u v) : list (sym2 V) := p.darts.map dart.edge @[simp] lemma support_cons {u v w : V} (h : G.adj u v) (p : G.walk v w) : (cons h p).support = u :: p.support := rfl +@[simp] lemma support_concat {u v w : V} (p : G.walk u v) (h : G.adj v w) : + (p.concat h).support = p.support.concat w := by induction p; simp [*, concat_nil] + @[simp] lemma support_copy {u v u' v'} (p : G.walk u v) (hu : u = u') (hv : v = v') : (p.copy hu hv).support = p.support := by { subst_vars, refl } @@ -427,6 +543,9 @@ edges_subset_edge_set p h @[simp] lemma darts_cons {u v w : V} (h : G.adj u v) (p : G.walk v w) : (cons h p).darts = ⟨(u, v), h⟩ :: p.darts := rfl +@[simp] lemma darts_concat {u v w : V} (p : G.walk u v) (h : G.adj v w) : + (p.concat h).darts = p.darts.concat ⟨(v, w), h⟩ := by induction p; simp [*, concat_nil] + @[simp] lemma darts_copy {u v u' v'} (p : G.walk u v) (hu : u = u') (hv : v = v') : (p.copy hu hv).darts = p.darts := by { subst_vars, refl } @@ -463,6 +582,9 @@ by simpa! using congr_arg list.init (map_fst_darts_append p) @[simp] lemma edges_cons {u v w : V} (h : G.adj u v) (p : G.walk v w) : (cons h p).edges = ⟦(u, v)⟧ :: p.edges := rfl +@[simp] lemma edges_concat {u v w : V} (p : G.walk u v) (h : G.adj v w) : + (p.concat h).edges = p.edges.concat ⟦(v, w)⟧ := by simp [edges] + @[simp] lemma edges_copy {u v u' v'} (p : G.walk u v) (hu : u = u') (hv : v = v') : (p.copy hu hv).edges = p.edges := by { subst_vars, refl } diff --git a/src/logic/basic.lean b/src/logic/basic.lean index f834107a1895c..2d51e8e082c71 100644 --- a/src/logic/basic.lean +++ b/src/logic/basic.lean @@ -908,9 +908,17 @@ lemma heq_of_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⟩ -lemma rec_heq_of_heq {β} {C : α → Sort*} {x : C a} {y : β} (eq : a = b) (h : x == y) : - @eq.rec α a C x b eq == y := -by subst eq; exact h +lemma rec_heq_of_heq {β} {C : α → Sort*} {x : C a} {y : β} (e : a = b) (h : x == y) : + @eq.rec α a C x b e == y := +by subst e; exact h + +lemma rec_heq_iff_heq {β} {C : α → Sort*} {x : C a} {y : β} {e : a = b} : + @eq.rec α a C x b e == y ↔ x == y := +by subst e + +lemma heq_rec_iff_heq {β} {C : α → Sort*} {x : β} {y : C a} {e : a = b} : + x == @eq.rec α a C y b e ↔ x == y := +by subst e protected lemma eq.congr {x₁ x₂ y₁ y₂ : α} (h₁ : x₁ = y₁) (h₂ : x₂ = y₂) : (x₁ = x₂) ↔ (y₁ = y₂) := From 2bbc7e3884ba234309d2a43b19144105a753292e Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 16 Feb 2023 17:18:56 +0000 Subject: [PATCH 0484/1291] feat(algebra/algebra/operations): the semiring of submodules over an algebra is idempotent (#18400) The same results are repeated for `ideal`. The two modifications of ported files are docstrings only, and are forward-ported at https://github.com/leanprover-community/mathlib4/pull/2175. Since this adjusts the import hierarchy slightly, there are some files which now need to qualify names that were previously unambiguous. One proof timed out, but the timeout went away after converting a messy term into tactic mode. --- src/algebra/algebra/operations.lean | 13 ++++++++----- .../group_action/sub_mul_action/pointwise.lean | 2 +- src/group_theory/nilpotent.lean | 4 ++-- src/group_theory/submonoid/pointwise.lean | 2 +- src/linear_algebra/adic_completion.lean | 2 +- src/linear_algebra/dimension.lean | 4 ++-- src/linear_algebra/dual.lean | 2 +- src/linear_algebra/matrix/diagonal.lean | 2 +- src/linear_algebra/matrix/to_lin.lean | 6 +++--- src/ring_theory/coprime/ideal.lean | 8 +++++--- src/ring_theory/dedekind_domain/ideal.lean | 2 +- src/ring_theory/ideal/operations.lean | 2 +- 12 files changed, 27 insertions(+), 22 deletions(-) diff --git a/src/algebra/algebra/operations.lean b/src/algebra/algebra/operations.lean index 5ed6f9175320f..92ce5700ebc6a 100644 --- a/src/algebra/algebra/operations.lean +++ b/src/algebra/algebra/operations.lean @@ -8,6 +8,7 @@ import algebra.algebra.equiv import algebra.module.submodule.pointwise import algebra.module.submodule.bilinear import algebra.module.opposites +import algebra.order.kleene import data.finset.pointwise import data.set.semiring import data.set.pointwise.big_operators @@ -305,8 +306,8 @@ by { simp_rw [(*), map₂_span_singleton_eq_map, exists_prop], refl } lemma mem_mul_span_singleton {x y : A} : x ∈ P * span R {y} ↔ ∃ z ∈ P, z * y = x := by { simp_rw [(*), map₂_span_singleton_eq_map_flip, exists_prop], refl } -/-- Sub-R-modules of an R-algebra form a semiring. -/ -instance : semiring (submodule R A) := +/-- Sub-R-modules of an R-algebra form an idempotent semiring. -/ +instance : idem_semiring (submodule R A) := { one_mul := submodule.one_mul, mul_one := submodule.mul_one, zero_mul := bot_mul, @@ -317,7 +318,9 @@ instance : semiring (submodule R A) := ..add_monoid_with_one.unary, ..submodule.pointwise_add_comm_monoid, ..submodule.has_one, - ..submodule.has_mul } + ..submodule.has_mul, + ..(by apply_instance : order_bot (submodule R A)), + ..(by apply_instance : lattice (submodule R A)) } variables (M) @@ -498,9 +501,9 @@ le_antisymm (mul_le.2 $ λ r hrm s hsn, mul_mem_mul_rev hsn hrm) (mul_le.2 $ λ r hrn s hsm, mul_mem_mul_rev hsm hrn) /-- Sub-R-modules of an R-algebra A form a semiring. -/ -instance : comm_semiring (submodule R A) := +instance : idem_comm_semiring (submodule R A) := { mul_comm := submodule.mul_comm, - .. submodule.semiring } + .. submodule.idem_semiring } lemma prod_span {ι : Type*} (s : finset ι) (M : ι → set A) : (∏ i in s, submodule.span R (M i)) = submodule.span R (∏ i in s, M i) := diff --git a/src/group_theory/group_action/sub_mul_action/pointwise.lean b/src/group_theory/group_action/sub_mul_action/pointwise.lean index e54d5c31913a0..d09c82d903019 100644 --- a/src/group_theory/group_action/sub_mul_action/pointwise.lean +++ b/src/group_theory/group_action/sub_mul_action/pointwise.lean @@ -15,7 +15,7 @@ import group_theory.group_action.sub_mul_action This file provides `sub_mul_action.monoid` and weaker typeclasses, which show that `sub_mul_action`s inherit the same pointwise multiplications as sets. -To match `submodule.semiring`, we do not put these in the `pointwise` locale. +To match `submodule.idem_semiring`, we do not put these in the `pointwise` locale. -/ diff --git a/src/group_theory/nilpotent.lean b/src/group_theory/nilpotent.lean index 3acfd6222da00..bea2a8d771b38 100644 --- a/src/group_theory/nilpotent.lean +++ b/src/group_theory/nilpotent.lean @@ -513,7 +513,7 @@ begin rw nilpotent_iff_lower_central_series at *, rcases hH with ⟨n, hn⟩, use (n + 1), - refine lower_central_series_succ_eq_bot (le_trans ((map_eq_bot_iff _).mp _) hf1), + refine lower_central_series_succ_eq_bot (le_trans ((subgroup.map_eq_bot_iff _).mp _) hf1), exact eq_bot_iff.mpr (hn ▸ (lower_central_series.map f n)), end @@ -524,7 +524,7 @@ lemma nilpotency_class_le_of_ker_le_center {H : Type*} [group H] (f : G →* H) begin rw ← lower_central_series_length_eq_nilpotency_class, apply nat.find_min', - refine lower_central_series_succ_eq_bot (le_trans ((map_eq_bot_iff _).mp _) hf1), + refine lower_central_series_succ_eq_bot (le_trans ((subgroup.map_eq_bot_iff _).mp _) hf1), apply eq_bot_iff.mpr, apply (le_trans (lower_central_series.map f _)), simp only [lower_central_series_nilpotency_class, le_bot_iff], diff --git a/src/group_theory/submonoid/pointwise.lean b/src/group_theory/submonoid/pointwise.lean index 599da656d223b..db1c25a50e117 100644 --- a/src/group_theory/submonoid/pointwise.lean +++ b/src/group_theory/submonoid/pointwise.lean @@ -32,7 +32,7 @@ Additionally, it provides various degrees of monoid structure: * `add_submonoid.mul_one_class` * `add_submonoid.semigroup` * `add_submonoid.monoid` -which is available globally to match the monoid structure implied by `submodule.semiring`. +which is available globally to match the monoid structure implied by `submodule.idem_semiring`. ## Implementation notes diff --git a/src/linear_algebra/adic_completion.lean b/src/linear_algebra/adic_completion.lean index 7548f11f3764d..bc42c89538b27 100644 --- a/src/linear_algebra/adic_completion.lean +++ b/src/linear_algebra/adic_completion.lean @@ -117,7 +117,7 @@ instance : is_Hausdorff I (Hausdorffification I M) := ⟨λ x, quotient.induction_on' x $ λ x hx, (quotient.mk_eq_zero _).2 $ (mem_infi _).2 $ λ n, begin have := comap_map_mkq (⨅ n : ℕ, I ^ n • ⊤ : submodule R M) (I ^ n • ⊤), simp only [sup_of_le_right (infi_le (λ n, (I ^ n • ⊤ : submodule R M)) n)] at this, - rw [← this, map_smul'', mem_comap, map_top, range_mkq, ← smodeq.zero], exact hx n + rw [← this, map_smul'', mem_comap, submodule.map_top, range_mkq, ← smodeq.zero], exact hx n end⟩ variables {M} [h : is_Hausdorff I N] diff --git a/src/linear_algebra/dimension.lean b/src/linear_algebra/dimension.lean index 7fcadfd6b57f5..3b72dc742207d 100644 --- a/src/linear_algebra/dimension.lean +++ b/src/linear_algebra/dimension.lean @@ -1063,7 +1063,7 @@ begin exact linear_map.ext_iff.1 eq c } }, { rw [← ker_eq_bot, ker_cod_restrict, ker_prod, hgd, bot_inf_eq] }, { rw [← range_eq_top, eq_top_iff, range_cod_restrict, ← map_le_iff_le_comap, - map_top, range_subtype], + submodule.map_top, range_subtype], rintros ⟨d, e⟩, have h := eq₂ d (-e), simp only [add_eq_zero_iff_eq_neg, linear_map.prod_apply, mem_ker, set_like.mem_coe, @@ -1079,7 +1079,7 @@ lemma dim_sup_add_dim_inf_eq (s t : submodule K V) : module.rank K s + module.rank K t := dim_add_dim_split (of_le le_sup_left) (of_le le_sup_right) (of_le inf_le_left) (of_le inf_le_right) begin - rw [← map_le_map_iff' (ker_subtype $ s ⊔ t), map_sup, map_top, + rw [← map_le_map_iff' (ker_subtype $ s ⊔ t), submodule.map_sup, submodule.map_top, ← linear_map.range_comp, ← linear_map.range_comp, subtype_comp_of_le, subtype_comp_of_le, range_subtype, range_subtype, range_subtype], exact le_rfl diff --git a/src/linear_algebra/dual.lean b/src/linear_algebra/dual.lean index 774b9fb9ad784..02d8548d4b701 100644 --- a/src/linear_algebra/dual.lean +++ b/src/linear_algebra/dual.lean @@ -394,7 +394,7 @@ lemma eval_range {ι : Type*} [_root_.finite ι] (b : basis ι R M) : (eval R M) begin classical, casesI nonempty_fintype ι, - rw [← b.to_dual_to_dual, range_comp, b.to_dual_range, map_top, to_dual_range _], + rw [← b.to_dual_to_dual, range_comp, b.to_dual_range, submodule.map_top, to_dual_range _], apply_instance end diff --git a/src/linear_algebra/matrix/diagonal.lean b/src/linear_algebra/matrix/diagonal.lean index 4adf3a8c42c18..74d5e149b0176 100644 --- a/src/linear_algebra/matrix/diagonal.lean +++ b/src/linear_algebra/matrix/diagonal.lean @@ -65,7 +65,7 @@ lemma range_diagonal [decidable_eq m] (w : m → K) : (diagonal w).to_lin'.range = (⨆ i ∈ {i | w i ≠ 0}, (linear_map.std_basis K (λi, K) i).range) := begin dsimp only [mem_set_of_eq], - rw [← map_top, ← supr_range_std_basis, map_supr], + rw [← submodule.map_top, ← supr_range_std_basis, submodule.map_supr], congr, funext i, rw [← linear_map.range_comp, diagonal_comp_std_basis, ← range_smul'] end diff --git a/src/linear_algebra/matrix/to_lin.lean b/src/linear_algebra/matrix/to_lin.lean index 7d5c8306db834..06025cd606198 100644 --- a/src/linear_algebra/matrix/to_lin.lean +++ b/src/linear_algebra/matrix/to_lin.lean @@ -270,9 +270,9 @@ lemma matrix.ker_to_lin'_eq_bot_iff {M : matrix n n R} : by simp only [submodule.eq_bot_iff, linear_map.mem_ker, matrix.to_lin'_apply] lemma matrix.range_to_lin' (M : matrix m n R) : M.to_lin'.range = span R (range Mᵀ) := -by simp_rw [range_eq_map, ←supr_range_std_basis, map_supr, range_eq_map, ←ideal.span_singleton_one, - ideal.span, submodule.map_span, image_image, image_singleton, matrix.to_lin'_apply, - M.mul_vec_std_basis_apply, supr_span, range_eq_Union] +by simp_rw [range_eq_map, ←supr_range_std_basis, submodule.map_supr, range_eq_map, + ←ideal.span_singleton_one, ideal.span, submodule.map_span, image_image, image_singleton, + matrix.to_lin'_apply, M.mul_vec_std_basis_apply, supr_span, range_eq_Union] /-- If `M` and `M'` are each other's inverse matrices, they provide an equivalence between `m → A` and `n → A` corresponding to `M.mul_vec` and `M'.mul_vec`. -/ diff --git a/src/ring_theory/coprime/ideal.lean b/src/ring_theory/coprime/ideal.lean index aa22abce01aca..9e27810bf08ca 100644 --- a/src/ring_theory/coprime/ideal.lean +++ b/src/ring_theory/coprime/ideal.lean @@ -62,9 +62,11 @@ begin apply this _ (finset.mem_cons_self _ _), rintro rfl, exact hat hx }, { have := submodule.coe_mem (μ a), simp only [mem_infi] at this, exact this _ (finset.subset_cons _ hb) ab.symm } } }, - { rintro ⟨hs, Hb⟩, obtain ⟨μ, hμ⟩ := ih.mpr hs, - obtain ⟨u, hu, v, hv, huv⟩ := submodule.mem_sup.mp - ((eq_top_iff_one _).mp $ sup_infi_eq_top $ λ b hb, Hb b hb $ by { rintro rfl, exact hat hb }), + { rintro ⟨hs, Hb⟩, + obtain ⟨μ, hμ⟩ := ih.mpr hs, + have := sup_infi_eq_top (λ b hb, Hb b hb (ne_of_mem_of_not_mem hb hat).symm), + rw [eq_top_iff_one, submodule.mem_sup] at this, + obtain ⟨u, hu, v, hv, huv⟩ := this, refine ⟨λ i, if hi : i = a then ⟨v, _⟩ else ⟨u * μ i, _⟩, _⟩, { simp only [mem_infi] at hv ⊢, intros j hj ij, rw [finset.mem_cons, ← hi] at hj, diff --git a/src/ring_theory/dedekind_domain/ideal.lean b/src/ring_theory/dedekind_domain/ideal.lean index 0c67434479528..53da5873c2207 100644 --- a/src/ring_theory/dedekind_domain/ideal.lean +++ b/src/ring_theory/dedekind_domain/ideal.lean @@ -613,7 +613,7 @@ instance fractional_ideal.cancel_comm_monoid_with_zero : instance ideal.cancel_comm_monoid_with_zero : cancel_comm_monoid_with_zero (ideal A) := -{ .. ideal.comm_semiring, +{ .. ideal.idem_comm_semiring, .. function.injective.cancel_comm_monoid_with_zero (coe_ideal_hom A⁰ (fraction_ring A)) coe_ideal_injective (ring_hom.map_zero _) (ring_hom.map_one _) (ring_hom.map_mul _) (ring_hom.map_pow _) } diff --git a/src/ring_theory/ideal/operations.lean b/src/ring_theory/ideal/operations.lean index f61bf2be7d13b..822a3ef92eeab 100644 --- a/src/ring_theory/ideal/operations.lean +++ b/src/ring_theory/ideal/operations.lean @@ -783,7 +783,7 @@ lemma is_radical_bot_of_no_zero_divisors {R} [comm_semiring R] [no_zero_divisors radical (⊥ : ideal R) = ⊥ := eq_bot_iff.2 is_radical_bot_of_no_zero_divisors -instance : comm_semiring (ideal R) := submodule.comm_semiring +instance : idem_comm_semiring (ideal R) := submodule.idem_comm_semiring variables (R) theorem top_pow (n : ℕ) : (⊤ ^ n : ideal R) = ⊤ := From 10d887272d1a72b99da88bcb301d1da9d9d33696 Mon Sep 17 00:00:00 2001 From: Michael Stoll Date: Thu, 16 Feb 2023 17:18:57 +0000 Subject: [PATCH 0485/1291] feat(algebra/continued_fractions/*): add some API lemmas (#18434) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR adds some API lemmas to the continued fractions files, leading to the recurrence ```lean lemma convergents_succ (v : K) (n : ℕ) : (of v).convergents (n + 1) = ⌊v⌋ + 1 / (of (int.fract v)⁻¹).convergents n ``` for the convergents of the continued fraction expansion of an element v of a linearly ordered floor field K. (Here `of` is `generalized_continued_fraction.of`.) This recurrence will then be used (in a later PR) for a proof of a theorem due to Legendre, which says that when x is a real number and p/q is a fraction such that |x - p/q| < 1/(2*q^2), then p/q is a convergent of the continued fraction expansion of x. See [this thread on Zulip](https://leanprover.zulipchat.com/#narrow/stream/116395-maths/topic/Diophantine.20approximation/near/327440641). --- .../approximation_corollaries.lean | 8 + .../computation/basic.lean | 3 +- .../computation/translations.lean | 142 ++++++++++++++++++ .../convergents_equiv.lean | 1 + .../continued_fractions/translations.lean | 8 + 5 files changed, 160 insertions(+), 2 deletions(-) diff --git a/src/algebra/continued_fractions/computation/approximation_corollaries.lean b/src/algebra/continued_fractions/computation/approximation_corollaries.lean index bf0999450af61..01f15e9e1a730 100644 --- a/src/algebra/continued_fractions/computation/approximation_corollaries.lean +++ b/src/algebra/continued_fractions/computation/approximation_corollaries.lean @@ -66,6 +66,14 @@ lemma of_convergents_eq_convergents' : (of v).convergents = (of v).convergents' := @continued_fraction.convergents_eq_convergents' _ _ (continued_fraction.of v) +/-- +The recurrence relation for the `convergents` of the continued fraction expansion +of an element `v` of `K` in terms of the convergents of the inverse of its fractional part. +-/ +lemma convergents_succ (n : ℕ): + (of v).convergents (n + 1) = ⌊v⌋ + 1 / (of (int.fract v)⁻¹).convergents n := +by rw [of_convergents_eq_convergents', convergents'_succ, of_convergents_eq_convergents'] + section convergence /-! ### Convergence diff --git a/src/algebra/continued_fractions/computation/basic.lean b/src/algebra/continued_fractions/computation/basic.lean index 89945ec471bc8..61ac955115559 100644 --- a/src/algebra/continued_fractions/computation/basic.lean +++ b/src/algebra/continued_fractions/computation/basic.lean @@ -31,8 +31,7 @@ For an example, refer to `int_fract_pair.stream`. - `generalized_continued_fraction.int_fract_pair.stream`: computes the stream of integer and fractional parts of a given value as described in the summary. - `generalized_continued_fraction.of`: computes the generalised continued fraction of a value `v`. - In fact, it computes a regular continued fraction that terminates if and only if `v` is rational - (those proofs will be added in a future commit). + In fact, it computes a regular continued fraction that terminates if and only if `v` is rational. ## Implementation Notes diff --git a/src/algebra/continued_fractions/computation/translations.lean b/src/algebra/continued_fractions/computation/translations.lean index a7eeeac76d51a..3d999f84a89ae 100644 --- a/src/algebra/continued_fractions/computation/translations.lean +++ b/src/algebra/continued_fractions/computation/translations.lean @@ -50,6 +50,8 @@ Here we state some lemmas that give us inversion rules and recurrences for the c stream of integer and fractional parts of a value. -/ +lemma stream_zero (v : K) : int_fract_pair.stream v 0 = some (int_fract_pair.of v) := rfl + variable {n : ℕ} lemma stream_eq_none_of_fr_eq_zero {ifp_n : int_fract_pair K} @@ -83,6 +85,26 @@ lemma succ_nth_stream_eq_some_iff {ifp_succ_n : int_fract_pair K} : ∧ int_fract_pair.of ifp_n.fr⁻¹ = ifp_succ_n := by simp [int_fract_pair.stream, ite_eq_iff] +/-- +An easier to use version of one direction of +`generalized_continued_fraction.int_fract_pair.succ_nth_stream_eq_some_iff`. +-/ +lemma stream_succ_of_some {p : int_fract_pair K} + (h : int_fract_pair.stream v n = some p) (h' : p.fr ≠ 0) : + int_fract_pair.stream v (n + 1) = some (int_fract_pair.of (p.fr)⁻¹) := +succ_nth_stream_eq_some_iff.mpr ⟨p, h, h', rfl⟩ + +/-- +The stream of `int_fract_pair`s of an integer stops after the first term. +-/ +lemma stream_succ_of_int (a : ℤ) (n : ℕ) : int_fract_pair.stream (a : K) (n + 1) = none := +begin + induction n with n ih, + { refine int_fract_pair.stream_eq_none_of_fr_eq_zero (int_fract_pair.stream_zero (a : K)) _, + simp only [int_fract_pair.of, int.fract_int_cast], }, + { exact int_fract_pair.succ_nth_stream_eq_none_iff.mpr (or.inl ih), } +end + lemma exists_succ_nth_stream_of_fr_zero {ifp_succ_n : int_fract_pair K} (stream_succ_nth_eq : int_fract_pair.stream v (n + 1) = some ifp_succ_n) (succ_nth_fr_eq_zero : ifp_succ_n.fr = 0) : @@ -96,6 +118,28 @@ begin simpa only [int_fract_pair.of, int.fract, sub_eq_zero] using succ_nth_fr_eq_zero end +/-- +A recurrence relation that expresses the `(n+1)`th term of the stream of `int_fract_pair`s +of `v` for non-integer `v` in terms of the `n`th term of the stream associated to +the inverse of the fractional part of `v`. +-/ +lemma stream_succ (h : int.fract v ≠ 0) (n : ℕ) : + int_fract_pair.stream v (n + 1) = int_fract_pair.stream (int.fract v)⁻¹ n := +begin + induction n with n ih, + { have H : (int_fract_pair.of v).fr = int.fract v := rfl, + rw [stream_zero, stream_succ_of_some (stream_zero v) (ne_of_eq_of_ne H h), H], }, + { cases eq_or_ne (int_fract_pair.stream (int.fract v)⁻¹ n) none with hnone hsome, + { rw hnone at ih, + rw [succ_nth_stream_eq_none_iff.mpr (or.inl hnone), + succ_nth_stream_eq_none_iff.mpr (or.inl ih)], }, + { obtain ⟨p, hp⟩ := option.ne_none_iff_exists'.mp hsome, + rw hp at ih, + cases eq_or_ne p.fr 0 with hz hnz, + { rw [stream_eq_none_of_fr_eq_zero hp hz, stream_eq_none_of_fr_eq_zero ih hz], }, + { rw [stream_succ_of_some hp hnz, stream_succ_of_some ih hnz], } } } +end + end int_fract_pair section head @@ -199,6 +243,104 @@ have int_fract_pair.stream v (n + 1) = some (int_fract_pair.of ifp_n.fr⁻¹), b { cases ifp_n, simp [int_fract_pair.stream, stream_nth_eq, nth_fr_ne_zero] }, nth_of_eq_some_of_succ_nth_int_fract_pair_stream this +open int int_fract_pair + +lemma of_s_head_aux (v : K) : + (of v).s.nth 0 = (int_fract_pair.stream v 1).bind (some ∘ λ p, {a := 1, b := p.b}) := +begin + rw [of, int_fract_pair.seq1, of._match_1], + simp only [seq.map_tail, seq.map, seq.tail, seq.head, seq.nth, stream.map], + rw [← stream.nth_succ, stream.nth, option.map], +end + +/-- +This gives the first pair of coefficients of the continued fraction of a non-integer `v`. +-/ +lemma of_s_head (h : fract v ≠ 0) : (of v).s.head = some ⟨1, ⌊(fract v)⁻¹⌋⟩ := +begin + change (of v).s.nth 0 = _, + rw [of_s_head_aux, stream_succ_of_some (stream_zero v) h, option.bind], + refl, +end + +variables (K) + +/-- +If `a` is an integer, then the coefficient sequence of its continued fraction is empty. +-/ +lemma of_s_of_int (a : ℤ) : (of (a : K)).s = seq.nil := +begin + have h : ∀ n, (of (a : K)).s.nth n = none, + { intro n, + induction n with n ih, + { rw [of_s_head_aux, stream_succ_of_int, option.bind], }, + { exact (of (a : K)).s.prop ih, } }, + exact seq.ext (λ n, (h n).trans (seq.nth_nil n).symm), +end + +variables {K} (v) + +/-- +Recurrence for the `generalized_continued_fraction.of` an element `v` of `K` in terms of +that of the inverse of the fractional part of `v`. +-/ +lemma of_s_succ (n : ℕ) : (of v).s.nth (n + 1) = (of (fract v)⁻¹).s.nth n := +begin + cases eq_or_ne (fract v) 0 with h h, + { obtain ⟨a, rfl⟩ : ∃ a : ℤ, v = a := ⟨⌊v⌋, eq_of_sub_eq_zero h⟩, + rw [fract_int_cast, inv_zero, of_s_of_int, ← cast_zero, of_s_of_int, seq.nth_nil, + seq.nth_nil], }, + cases eq_or_ne ((of (fract v)⁻¹).s.nth n) none with h₁ h₁, + { rwa [h₁, ← terminated_at_iff_s_none, + of_terminated_at_n_iff_succ_nth_int_fract_pair_stream_eq_none, stream_succ h, + ← of_terminated_at_n_iff_succ_nth_int_fract_pair_stream_eq_none, + terminated_at_iff_s_none], }, + { obtain ⟨p, hp⟩ := option.ne_none_iff_exists'.mp h₁, + obtain ⟨p', hp'₁, _⟩ := exists_succ_nth_stream_of_gcf_of_nth_eq_some hp, + have Hp := nth_of_eq_some_of_succ_nth_int_fract_pair_stream hp'₁, + rw [← stream_succ h] at hp'₁, + rw [Hp, nth_of_eq_some_of_succ_nth_int_fract_pair_stream hp'₁], } +end + +/-- +This expresses the tail of the coefficient sequence of the `generalized_continued_fraction.of` +an element `v` of `K` as the coefficient sequence of that of the inverse of the +fractional part of `v`. +-/ +lemma of_s_tail : (of v).s.tail = (of (fract v)⁻¹).s := +seq.ext $ λ n, seq.nth_tail (of v).s n ▸ of_s_succ v n + +variables (K) (n) + +/-- +If `a` is an integer, then the `convergents'` of its continued fraction expansion +are all equal to `a`. +-/ +lemma convergents'_of_int (a : ℤ) : (of (a : K)).convergents' n = a := +begin + induction n with n ih, + { simp only [zeroth_convergent'_eq_h, of_h_eq_floor, floor_int_cast], }, + { rw [convergents', of_h_eq_floor, floor_int_cast, add_right_eq_self], + exact convergents'_aux_succ_none ((of_s_of_int K a).symm ▸ seq.nth_nil 0) _, } +end + +variables {K} (v) + +/-- +The recurrence relation for the `convergents'` of the continued fraction expansion +of an element `v` of `K` in terms of the convergents of the inverse of its fractional part. +-/ +lemma convergents'_succ : + (of v).convergents' (n + 1) = ⌊v⌋ + 1 / (of (fract v)⁻¹).convergents' n := +begin + cases eq_or_ne (fract v) 0 with h h, + { obtain ⟨a, rfl⟩ : ∃ a : ℤ, v = a := ⟨⌊v⌋, eq_of_sub_eq_zero h⟩, + rw [convergents'_of_int, fract_int_cast, inv_zero, ← cast_zero, + convergents'_of_int, cast_zero, div_zero, add_zero, floor_int_cast], }, + { rw [convergents', of_h_eq_floor, add_right_inj, convergents'_aux_succ_some (of_s_head h)], + exact congr_arg ((/) 1) (by rw [convergents', of_h_eq_floor, add_right_inj, of_s_tail]), } +end + end values end sequence end generalized_continued_fraction diff --git a/src/algebra/continued_fractions/convergents_equiv.lean b/src/algebra/continued_fractions/convergents_equiv.lean index dee99ac682d10..439bf0d209637 100644 --- a/src/algebra/continued_fractions/convergents_equiv.lean +++ b/src/algebra/continued_fractions/convergents_equiv.lean @@ -18,6 +18,7 @@ direct evaluation (`convergents'`)) for `gcf`s on linear ordered fields. We foll [hardy2008introduction], Chapter 10. Here's a sketch: Let `c` be a continued fraction `[h; (a₀, b₀), (a₁, b₁), (a₂, b₂),...]`, visually: + a₀ h + --------------------------- a₁ diff --git a/src/algebra/continued_fractions/translations.lean b/src/algebra/continued_fractions/translations.lean index 5849efbafcf65..3a1f74dc5a8dc 100644 --- a/src/algebra/continued_fractions/translations.lean +++ b/src/algebra/continued_fractions/translations.lean @@ -119,5 +119,13 @@ lemma zeroth_convergent'_aux_eq_zero {s : seq $ pair K} : convergents'_aux s 0 = @[simp] lemma zeroth_convergent'_eq_h : g.convergents' 0 = g.h := by simp [convergents'] +lemma convergents'_aux_succ_none {s : seq (pair K)} (h : s.head = none) (n : ℕ) : + convergents'_aux s (n + 1) = 0 := +by rw [convergents'_aux, h, convergents'_aux._match_1] + +lemma convergents'_aux_succ_some {s : seq (pair K)} {p : pair K} (h : s.head = some p) (n : ℕ) : + convergents'_aux s (n + 1) = p.a / (p.b + convergents'_aux s.tail n) := +by rw [convergents'_aux, h, convergents'_aux._match_1] + end with_division_ring end generalized_continued_fraction From 5455cb0b5f3be8f8b22c55366fbc6e380dbff579 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 16 Feb 2023 17:18:58 +0000 Subject: [PATCH 0486/1291] chore(linear_algebra/dual): prove a lemma with rfl (#18444) I was surprised that `dsimp` didn't clean this up for me. The proof that used to be here was certainly not interesting. --- src/linear_algebra/dual.lean | 6 +----- 1 file changed, 1 insertion(+), 5 deletions(-) diff --git a/src/linear_algebra/dual.lean b/src/linear_algebra/dual.lean index 02d8548d4b701..0244ffc815048 100644 --- a/src/linear_algebra/dual.lean +++ b/src/linear_algebra/dual.lean @@ -118,11 +118,7 @@ instance : has_coe_to_fun (dual R M) (λ _, M → R) := ⟨linear_map.to_fun⟩ `module.eval_equiv`. -/ def eval : M →ₗ[R] (dual R (dual R M)) := linear_map.flip linear_map.id -@[simp] lemma eval_apply (v : M) (a : dual R M) : eval R M v a = a v := -begin - dunfold eval, - rw [linear_map.flip_apply, linear_map.id_apply] -end +@[simp] lemma eval_apply (v : M) (a : dual R M) : eval R M v a = a v := rfl variables {R M} {M' : Type*} [add_comm_monoid M'] [module R M'] From 740acc0e6f9adf4423f92a485d0456fc271482da Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Thu, 16 Feb 2023 19:50:52 +0000 Subject: [PATCH 0487/1291] =?UTF-8?q?feat(algebra/order):=20Unions=20of=20?= =?UTF-8?q?interval=20translates=20by=20=E2=84=A4=20(#18427)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Write a linearly ordered abelian group as a union of pairwise disjoint intervals `⋃ (n : ℤ), Ioc (a + n • b) (a + (n + 1) • b)` (in multiple variations). Split off from #18425, which in turn was split from #18392. --- src/algebra/order/to_interval_mod.lean | 74 ++++++++++++++++++ src/data/set/intervals/group.lean | 102 ++++++++++++++++++++++++- 2 files changed, 174 insertions(+), 2 deletions(-) diff --git a/src/algebra/order/to_interval_mod.lean b/src/algebra/order/to_interval_mod.lean index f696f550e5552..3b484635db0c3 100644 --- a/src/algebra/order/to_interval_mod.lean +++ b/src/algebra/order/to_interval_mod.lean @@ -733,3 +733,77 @@ lemma to_Ico_mod_zero_one (x : α) : to_Ico_mod (0 : α) zero_lt_one x = int.fra by simp [to_Ico_mod_eq_add_fract_mul] end linear_ordered_field + +/-! ### Lemmas about unions of translates of intervals -/ +section union + +open set int + +section linear_ordered_add_comm_group + +variables {α : Type*} [linear_ordered_add_comm_group α] [archimedean α] (a : α) {b : α} (hb : 0 < b) +include hb + +lemma Union_Ioc_add_zsmul : (⋃ (n : ℤ), Ioc (a + n • b) (a + (n + 1) • b)) = univ := +begin + refine eq_univ_iff_forall.mpr (λ x, mem_Union.mpr _), + rcases sub_to_Ioc_div_zsmul_mem_Ioc a hb x with ⟨hl, hr⟩, + refine ⟨to_Ioc_div a hb x, ⟨lt_sub_iff_add_lt.mp hl, _⟩⟩, + rw [add_smul, one_smul, ←add_assoc], + convert sub_le_iff_le_add.mp hr using 1, abel, +end + +lemma Union_Ico_add_zsmul : (⋃ (n : ℤ), Ico (a + n • b) (a + (n + 1) • b)) = univ := +begin + refine eq_univ_iff_forall.mpr (λ x, mem_Union.mpr _), + rcases sub_to_Ico_div_zsmul_mem_Ico a hb x with ⟨hl, hr⟩, + refine ⟨to_Ico_div a hb x, ⟨le_sub_iff_add_le.mp hl, _⟩⟩, + rw [add_smul, one_smul, ←add_assoc], + convert sub_lt_iff_lt_add.mp hr using 1, abel, +end + +lemma Union_Icc_add_zsmul : (⋃ (n : ℤ), Icc (a + n • b) (a + (n + 1) • b)) = univ := +by simpa only [Union_Ioc_add_zsmul a hb, univ_subset_iff] using + Union_mono (λ n : ℤ, (Ioc_subset_Icc_self : Ioc (a + n • b) (a + (n + 1) • b) ⊆ Icc _ _)) + +lemma Union_Ioc_zsmul : (⋃ (n : ℤ), Ioc (n • b) ((n + 1) • b)) = univ := +by simpa only [zero_add] using Union_Ioc_add_zsmul 0 hb + +lemma Union_Ico_zsmul : (⋃ (n : ℤ), Ico (n • b) ((n + 1) • b)) = univ := +by simpa only [zero_add] using Union_Ico_add_zsmul 0 hb + +lemma Union_Icc_zsmul : (⋃ (n : ℤ), Icc (n • b) ((n + 1) • b)) = univ := +by simpa only [zero_add] using Union_Icc_add_zsmul 0 hb + +end linear_ordered_add_comm_group + +section linear_ordered_ring + +variables {α : Type*} [linear_ordered_ring α] [archimedean α] (a : α) + +lemma Union_Ioc_add_int_cast : (⋃ (n : ℤ), Ioc (a + n) (a + n + 1)) = set.univ := +by simpa only [zsmul_one, int.cast_add, int.cast_one, ←add_assoc] + using Union_Ioc_add_zsmul a zero_lt_one + +lemma Union_Ico_add_int_cast : (⋃ (n : ℤ), Ico (a + n) (a + n + 1)) = set.univ := +by simpa only [zsmul_one, int.cast_add, int.cast_one, ←add_assoc] + using Union_Ico_add_zsmul a zero_lt_one + +lemma Union_Icc_add_int_cast : (⋃ (n : ℤ), Icc (a + n) (a + n + 1)) = set.univ := +by simpa only [zsmul_one, int.cast_add, int.cast_one, ←add_assoc] + using Union_Icc_add_zsmul a (zero_lt_one' α) + +variables (α) + +lemma Union_Ioc_int_cast : (⋃ (n : ℤ), Ioc (n : α) (n + 1)) = set.univ := +by simpa only [zero_add] using Union_Ioc_add_int_cast (0 : α) + +lemma Union_Ico_int_cast : (⋃ (n : ℤ), Ico (n : α) (n + 1)) = set.univ := +by simpa only [zero_add] using Union_Ico_add_int_cast (0 : α) + +lemma Union_Icc_int_cast : (⋃ (n : ℤ), Icc (n : α) (n + 1)) = set.univ := +by simpa only [zero_add] using Union_Icc_add_int_cast (0 : α) + +end linear_ordered_ring + +end union diff --git a/src/data/set/intervals/group.lean b/src/data/set/intervals/group.lean index ad8805380b283..b72182b7ff23f 100644 --- a/src/data/set/intervals/group.lean +++ b/src/data/set/intervals/group.lean @@ -4,12 +4,15 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro, Patrick Massot, Yury Kudryashov, Rémy Degenne -/ import data.set.intervals.basic +import data.set.pairwise import algebra.order.group.abs +import algebra.group_power.lemmas -/-! ### Lemmas about arithmetic operations and intervals. +/-! ### Lemmas about arithmetic operations and intervals. > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> Any changes to this file require a corresponding PR to mathlib4.-/ +> Any changes to this file require a corresponding PR to mathlib4. +-/ variables {α : Type*} @@ -98,4 +101,99 @@ end end linear_ordered_add_comm_group +/-! ### Lemmas about disjointness of translates of intervals -/ +section pairwise_disjoint + +section ordered_comm_group + +variables [ordered_comm_group α] (a b : α) + +@[to_additive] +lemma pairwise_disjoint_Ioc_mul_zpow : + pairwise (disjoint on λ n : ℤ, Ioc (a * b ^ n) (a * b ^ (n + 1))) := +begin + simp_rw [function.on_fun, set.disjoint_iff], + intros m n hmn x hx, + apply hmn, + have hb : 1 < b, + { have : a * b ^ m < a * b ^ (m + 1), from hx.1.1.trans_le hx.1.2, + rwa [mul_lt_mul_iff_left, ←mul_one (b ^ m), zpow_add_one, mul_lt_mul_iff_left] at this }, + have i1 := hx.1.1.trans_le hx.2.2, + have i2 := hx.2.1.trans_le hx.1.2, + rw [mul_lt_mul_iff_left, zpow_lt_zpow_iff hb, int.lt_add_one_iff] at i1 i2, + exact le_antisymm i1 i2 +end + +@[to_additive] +lemma pairwise_disjoint_Ico_mul_zpow : + pairwise (disjoint on λ n : ℤ, Ico (a * b ^ n) (a * b ^ (n + 1))) := +begin + simp_rw [function.on_fun, set.disjoint_iff], + intros m n hmn x hx, + apply hmn, + have hb : 1 < b, + { have : a * b ^ m < a * b ^ (m + 1), from hx.1.1.trans_lt hx.1.2, + rwa [mul_lt_mul_iff_left, ←mul_one (b ^ m), zpow_add_one, mul_lt_mul_iff_left] at this }, + have i1 := hx.1.1.trans_lt hx.2.2, + have i2 := hx.2.1.trans_lt hx.1.2, + rw [mul_lt_mul_iff_left, zpow_lt_zpow_iff hb, int.lt_add_one_iff] at i1 i2, + exact le_antisymm i1 i2, +end + +@[to_additive] +lemma pairwise_disjoint_Ioo_mul_zpow : + pairwise (disjoint on λ n : ℤ, Ioo (a * b ^ n) (a * b ^ (n + 1))) := +λ m n hmn, (pairwise_disjoint_Ioc_mul_zpow a b hmn).mono Ioo_subset_Ioc_self Ioo_subset_Ioc_self + +@[to_additive] +lemma pairwise_disjoint_Ioc_zpow : + pairwise (disjoint on λ n : ℤ, Ioc (b ^ n) (b ^ (n + 1))) := +by simpa only [one_mul] using pairwise_disjoint_Ioc_mul_zpow 1 b + +@[to_additive] +lemma pairwise_disjoint_Ico_zpow : + pairwise (disjoint on λ n : ℤ, Ico (b ^ n) (b ^ (n + 1))) := +by simpa only [one_mul] using pairwise_disjoint_Ico_mul_zpow 1 b + +@[to_additive] +lemma pairwise_disjoint_Ioo_zpow : + pairwise (disjoint on λ n : ℤ, Ioo (b ^ n) (b ^ (n + 1))) := +by simpa only [one_mul] using pairwise_disjoint_Ioo_mul_zpow 1 b + +end ordered_comm_group + +section ordered_ring + +variables [ordered_ring α] (a : α) + +lemma pairwise_disjoint_Ioc_add_int_cast : + pairwise (disjoint on λ n : ℤ, Ioc (a + n) (a + n + 1)) := +by simpa only [zsmul_one, int.cast_add, int.cast_one, ←add_assoc] + using pairwise_disjoint_Ioc_add_zsmul a (1 : α) + +lemma pairwise_disjoint_Ico_add_int_cast : + pairwise (disjoint on λ n : ℤ, Ico (a + n) (a + n + 1)) := +by simpa only [zsmul_one, int.cast_add, int.cast_one, ←add_assoc] + using pairwise_disjoint_Ico_add_zsmul a (1 : α) + +lemma pairwise_disjoint_Ioo_add_int_cast : + pairwise (disjoint on λ n : ℤ, Ioo (a + n) (a + n + 1)) := +by simpa only [zsmul_one, int.cast_add, int.cast_one, ←add_assoc] + using pairwise_disjoint_Ioo_add_zsmul a (1 : α) + +variables (α) + +lemma pairwise_disjoint_Ico_int_cast : pairwise (disjoint on λ n : ℤ, Ico (n : α) (n + 1)) := +by simpa only [zero_add] using pairwise_disjoint_Ico_add_int_cast (0 : α) + +lemma pairwise_disjoint_Ioo_int_cast : pairwise (disjoint on λ n : ℤ, Ioo (n : α) (n + 1)) := +by simpa only [zero_add] using pairwise_disjoint_Ioo_add_int_cast (0 : α) + +lemma pairwise_disjoint_Ioc_int_cast : pairwise (disjoint on λ n : ℤ, Ioc (n : α) (n + 1)) := +by simpa only [zero_add] using pairwise_disjoint_Ioc_add_int_cast (0 : α) + +end ordered_ring + +end pairwise_disjoint + end set From 4330aae21f538b862f8aead371cfb6ee556398f1 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 16 Feb 2023 21:51:10 +0000 Subject: [PATCH 0488/1291] feat(topology/algebra/order/upper_lower): The closure of an upper set (#16975) Topological facts about order-connected sets. --- src/group_theory/group_action/opposite.lean | 1 + src/topology/algebra/order/upper_lower.lean | 93 +++++++++++++++++++++ 2 files changed, 94 insertions(+) create mode 100644 src/topology/algebra/order/upper_lower.lean diff --git a/src/group_theory/group_action/opposite.lean b/src/group_theory/group_action/opposite.lean index e66543bb1f8fc..6d3b9d2174b6e 100644 --- a/src/group_theory/group_action/opposite.lean +++ b/src/group_theory/group_action/opposite.lean @@ -104,6 +104,7 @@ instance mul_action.opposite_regular.is_pretransitive {G : Type*} [group G] : smul_comm_class α αᵐᵒᵖ α := smul_comm_class.symm _ _ _ +@[to_additive] instance comm_semigroup.is_central_scalar [comm_semigroup α] : is_central_scalar α α := ⟨λ r m, mul_comm _ _⟩ diff --git a/src/topology/algebra/order/upper_lower.lean b/src/topology/algebra/order/upper_lower.lean new file mode 100644 index 0000000000000..3fdb7e5714249 --- /dev/null +++ b/src/topology/algebra/order/upper_lower.lean @@ -0,0 +1,93 @@ +/- +Copyright (c) 2022 Yaël Dillies. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies +-/ +import algebra.order.upper_lower +import topology.algebra.group.basic + +/-! +# Topological facts about upper/lower/order-connected sets + +The topological closure and interior of an upper/lower/order-connected set is an +upper/lower/order-connected set (with the notable exception of the closure of an order-connected +set). + +## Notes + +The lemmas don't mention additive/multiplicative operations. As a result, we decide to prime the +multiplicative lemma names to indicate that there is probably a common generalisation to each pair +of additive/multiplicative lemma. +-/ + +open function set +open_locale pointwise + +/-- Ad hoc class stating that the closure of an upper set is an upper set. This is used to state +lemmas that do not mention algebraic operations for both the additive and multiplicative versions +simultaneously. If you find a satisfying replacement for this typeclass, please remove it! -/ +class has_upper_lower_closure (α : Type*) [topological_space α] [preorder α] : Prop := +(is_upper_set_closure : ∀ s : set α, is_upper_set s → is_upper_set (closure s)) +(is_lower_set_closure : ∀ s : set α, is_lower_set s → is_lower_set (closure s)) +(is_open_upper_closure : ∀ s : set α, is_open s → is_open (upper_closure s : set α)) +(is_open_lower_closure : ∀ s : set α, is_open s → is_open (lower_closure s : set α)) + +variables {α : Type*} [topological_space α] + +@[to_additive, priority 100] -- See note [lower instance priority] +instance ordered_comm_group.to_has_upper_lower_closure [ordered_comm_group α] + [has_continuous_const_smul α α] : has_upper_lower_closure α := +{ is_upper_set_closure := λ s h x y hxy hx, closure_mono (h.smul_subset $ one_le_div'.2 hxy) $ + by { rw closure_smul, exact ⟨x, hx, div_mul_cancel' _ _⟩ }, + is_lower_set_closure := λ s h x y hxy hx, closure_mono (h.smul_subset $ div_le_one'.2 hxy) $ + by { rw closure_smul, exact ⟨x, hx, div_mul_cancel' _ _⟩ }, + is_open_upper_closure := λ s hs, by { rw [←mul_one s, ←mul_upper_closure], exact hs.mul_right }, + is_open_lower_closure := λ s hs, by { rw [←mul_one s, ←mul_lower_closure], exact hs.mul_right } } + +variables [preorder α] [has_upper_lower_closure α] {s : set α} + +protected lemma is_upper_set.closure : is_upper_set s → is_upper_set (closure s) := +has_upper_lower_closure.is_upper_set_closure _ + +protected lemma is_lower_set.closure : is_lower_set s → is_lower_set (closure s) := +has_upper_lower_closure.is_lower_set_closure _ + +protected lemma is_open.upper_closure : is_open s → is_open (upper_closure s : set α) := +has_upper_lower_closure.is_open_upper_closure _ + +protected lemma is_open.lower_closure : is_open s → is_open (lower_closure s : set α) := +has_upper_lower_closure.is_open_lower_closure _ + +instance : has_upper_lower_closure αᵒᵈ := +{ is_upper_set_closure := @is_lower_set.closure α _ _ _, + is_lower_set_closure := @is_upper_set.closure α _ _ _, + is_open_upper_closure := @is_open.lower_closure α _ _ _, + is_open_lower_closure := @is_open.upper_closure α _ _ _ } + +/- +Note: `s.ord_connected` does not imply `(closure s).ord_connected`, as we can see by taking +`s := Ioo 0 1 × Ioo 1 2 ∪ Ioo 2 3 × Ioo 0 1` because then +`closure s = Icc 0 1 × Icc 1 2 ∪ Icc 2 3 × Icc 0 1` is not order-connected as +`(1, 1) ∈ closure s`, `(2, 1) ∈ closure s` but `Icc (1, 1) (2, 1) ⊈ closure s`. + +`s` looks like +``` +xxooooo +xxooooo +oooooxx +oooooxx +``` +-/ + +protected lemma is_upper_set.interior (h : is_upper_set s) : is_upper_set (interior s) := +by { rw [←is_lower_set_compl, ←closure_compl], exact h.compl.closure } + +protected lemma is_lower_set.interior (h : is_lower_set s) : is_lower_set (interior s) := +h.of_dual.interior + +protected lemma set.ord_connected.interior (h : s.ord_connected) : (interior s).ord_connected := +begin + rw [←h.upper_closure_inter_lower_closure, interior_inter], + exact (upper_closure s).upper.interior.ord_connected.inter + (lower_closure s).lower.interior.ord_connected, +end From 1df9a17426d2bd593e22234897989eb2ddb4a00d Mon Sep 17 00:00:00 2001 From: Alex J Best Date: Thu, 16 Feb 2023 21:51:11 +0000 Subject: [PATCH 0489/1291] fix(.github/workflows): try a fix for review bodies (#18455) cf. https://github.com/leanprover-community/mathlib/pull/18400#pullrequestreview-1301127093 as reviews can only be posted on pr's we don't need these checks, and they were stopping the actions from firing. --- .github/workflows/maintainer_merge_review.yml | 2 +- .github/workflows/maintainer_merge_review_comment.yml | 2 +- 2 files changed, 2 insertions(+), 2 deletions(-) diff --git a/.github/workflows/maintainer_merge_review.yml b/.github/workflows/maintainer_merge_review.yml index 60e48de15ff36..5003f2f902825 100644 --- a/.github/workflows/maintainer_merge_review.yml +++ b/.github/workflows/maintainer_merge_review.yml @@ -7,7 +7,7 @@ on: jobs: ping_zulip: name: Ping maintainers on Zulip - if: (github.event.issue.pull_request != 'null') && (startsWith(github.event.review.body, 'maintainer merge') || contains(toJSON(github.event.review.body), '\r\nmaintainer merge')) + if: (startsWith(github.event.review.body, 'maintainer merge') || contains(toJSON(github.event.review.body), '\r\nmaintainer merge')) runs-on: ubuntu-latest steps: - name: Check whether user is part of mathlib-reviewers team diff --git a/.github/workflows/maintainer_merge_review_comment.yml b/.github/workflows/maintainer_merge_review_comment.yml index d3cb0c7a7d0df..2f7f40443e5b4 100644 --- a/.github/workflows/maintainer_merge_review_comment.yml +++ b/.github/workflows/maintainer_merge_review_comment.yml @@ -7,7 +7,7 @@ on: jobs: ping_zulip: name: Ping maintainers on Zulip - if: (github.event.issue.pull_request != 'null') && (startsWith(github.event.comment.body, 'maintainer merge') || contains(toJSON(github.event.comment.body), '\r\nmaintainer merge')) + if: (startsWith(github.event.comment.body, 'maintainer merge') || contains(toJSON(github.event.comment.body), '\r\nmaintainer merge')) runs-on: ubuntu-latest steps: - name: Check whether user is part of mathlib-reviewers team From 973ed779517bcba5d1b4fb750348a7403529e8b0 Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Fri, 17 Feb 2023 00:41:39 +0000 Subject: [PATCH 0490/1291] ci(.github/workflows): lint for formatting issues in references.bib (#16614) Co-authored-by: Bryan Gin-ge Chen --- .github/workflows/bors.yml | 10 ++++- .github/workflows/build.yml | 10 ++++- .github/workflows/build.yml.in | 10 ++++- .github/workflows/build_fork.yml | 10 ++++- docs/references.bib | 76 +++++++++++++++++--------------- scripts/lint-bib.sh | 8 ++++ 6 files changed, 84 insertions(+), 40 deletions(-) create mode 100755 scripts/lint-bib.sh diff --git a/.github/workflows/bors.yml b/.github/workflows/bors.yml index e808ac75f1a2e..8b13cb25e519b 100644 --- a/.github/workflows/bors.yml +++ b/.github/workflows/bors.yml @@ -34,16 +34,24 @@ jobs: steps: - uses: actions/checkout@v2 + - name: Install bibtool + if: ${{ 'bors' == 'ubuntu-latest' }} + run: sudo apt install -y bibtool + - name: install Python if: ${{ 'bors' == 'ubuntu-latest' }} uses: actions/setup-python@v2 with: python-version: 3.8 - - name: lint + - name: lint style run: | ./scripts/lint-style.sh + - name: lint references.bib + run: | + ./scripts/lint-bib.sh + build: if: github.repository == 'leanprover-community/mathlib' name: Build mathlib diff --git a/.github/workflows/build.yml b/.github/workflows/build.yml index 0656b61d919ae..ad8edebc51b94 100644 --- a/.github/workflows/build.yml +++ b/.github/workflows/build.yml @@ -42,16 +42,24 @@ jobs: steps: - uses: actions/checkout@v2 + - name: Install bibtool + if: ${{ 'ubuntu-latest' == 'ubuntu-latest' }} + run: sudo apt install -y bibtool + - name: install Python if: ${{ 'ubuntu-latest' == 'ubuntu-latest' }} uses: actions/setup-python@v2 with: python-version: 3.8 - - name: lint + - name: lint style run: | ./scripts/lint-style.sh + - name: lint references.bib + run: | + ./scripts/lint-bib.sh + build: if: github.repository == 'leanprover-community/mathlib' name: Build mathlib diff --git a/.github/workflows/build.yml.in b/.github/workflows/build.yml.in index f17250cc6b119..b9ebf269ddafe 100644 --- a/.github/workflows/build.yml.in +++ b/.github/workflows/build.yml.in @@ -20,16 +20,24 @@ jobs: steps: - uses: actions/checkout@v2 + - name: Install bibtool + if: ${{ 'STYLE_LINT_RUNNER' == 'ubuntu-latest' }} + run: sudo apt install -y bibtool + - name: install Python if: ${{ 'STYLE_LINT_RUNNER' == 'ubuntu-latest' }} uses: actions/setup-python@v2 with: python-version: 3.8 - - name: lint + - name: lint style run: | ./scripts/lint-style.sh + - name: lint references.bib + run: | + ./scripts/lint-bib.sh + build: if: github.repository MAIN_OR_FORK 'leanprover-community/mathlib' name: Build mathlibJOB_NAME diff --git a/.github/workflows/build_fork.yml b/.github/workflows/build_fork.yml index 8f12c3bc1f37d..58d0322673660 100644 --- a/.github/workflows/build_fork.yml +++ b/.github/workflows/build_fork.yml @@ -40,16 +40,24 @@ jobs: steps: - uses: actions/checkout@v2 + - name: Install bibtool + if: ${{ 'ubuntu-latest' == 'ubuntu-latest' }} + run: sudo apt install -y bibtool + - name: install Python if: ${{ 'ubuntu-latest' == 'ubuntu-latest' }} uses: actions/setup-python@v2 with: python-version: 3.8 - - name: lint + - name: lint style run: | ./scripts/lint-style.sh + - name: lint references.bib + run: | + ./scripts/lint-bib.sh + build: if: github.repository != 'leanprover-community/mathlib' name: Build mathlib (fork) diff --git a/docs/references.bib b/docs/references.bib index d5afefcaf7e02..acbfb060539cc 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -226,8 +226,8 @@ @Article{ birkhoff1942 @Book{ bollobas1986, author = {Bollob\'{a}s, B\'{e}la}, - title = {Combinatorics: Set Systems, Hypergraphs, Families of Vectors, and Combinatorial - Probability}, + title = {Combinatorics: Set Systems, Hypergraphs, Families of + Vectors, and Combinatorial Probability}, year = {1986}, isbn = {0521330599}, publisher = {Cambridge University Press} @@ -477,6 +477,23 @@ @InProceedings{ carneiro2019 bibsource = {dblp computer science bibliography, https://dblp.org} } +@Article{ cassels1950, + author = {Cassels, J. W. S.}, + title = {Some metrical theorems in {D}iophantine approximation. + {I}}, + journal = {Proc. Cambridge Philos. Soc.}, + fjournal = {Proceedings of the Cambridge Philosophical Society}, + volume = {46}, + year = {1950}, + pages = {209--218}, + issn = {0008-1981}, + mrclass = {10.0X}, + mrnumber = {36787}, + mrreviewer = {P. Erd\H{o}s}, + doi = {10.1017/s0305004100025676}, + url = {https://doi.org/10.1017/s0305004100025676} +} + @Book{ cassels1967algebraic, title = {Algebraic number theory}, author = {Cassels, John William Scott and Fr{\"o}hlich, Albrecht}, @@ -492,22 +509,6 @@ @Book{ cassels1967algebraic mrclass = {00.04 (10.00)} } -@Article{ cassels1950, - author = {Cassels, J. W. S.}, - title = {Some metrical theorems in {D}iophantine approximation. {I}}, - journal = {Proc. Cambridge Philos. Soc.}, - fjournal = {Proceedings of the Cambridge Philosophical Society}, - volume = {46}, - year = {1950}, - pages = {209--218}, - issn = {0008-1981}, - mrclass = {10.0X}, - mrnumber = {36787}, - mrreviewer = {P. Erd\H{o}s}, - doi = {10.1017/s0305004100025676}, - url = {https://doi.org/10.1017/s0305004100025676}, -} - @InProceedings{ Chou1994, author = {Chou, Ching-Tsun}, booktitle = {Higher Order Logic Theorem Proving and Its Applications}, @@ -921,7 +922,7 @@ @Article{ Gallagher1961 mrnumber = {133297}, mrreviewer = {J. W. S. Cassels}, doi = {10.2969/jmsj/01340342}, - url = {https://doi.org/10.2969/jmsj/01340342}, + url = {https://doi.org/10.2969/jmsj/01340342} } @InProceedings{ Gallier2011Notes, @@ -1388,7 +1389,7 @@ @Article{ kleitman1966 zbl = {0141.00801} } -@Article{ KoukoulopoulosMaynard2020, +@Article{ KoukoulopoulosMaynard2020, author = {Koukoulopoulos, Dimitris and Maynard, James}, title = {On the {D}uffin-{S}chaeffer conjecture}, journal = {Ann. of Math. (2)}, @@ -1402,23 +1403,26 @@ @Article{ KoukoulopoulosMaynard2020 mrnumber = {4125453}, mrreviewer = {Sam Chow}, doi = {10.4007/annals.2020.192.1.5}, - url = {https://doi.org/10.4007/annals.2020.192.1.5}, + url = {https://doi.org/10.4007/annals.2020.192.1.5} } -@Article{ kozen1994, - title = {A Completeness Theorem for Kleene Algebras and the Algebra of Regular Events}, - journal = {Information and Computation}, - volume = {110}, - number = {2}, - pages = {366-390}, - year = {1994}, - issn = {0890-5401}, - doi = {https://doi.org/10.1006/inco.1994.1037}, - url = {https://www.sciencedirect.com/science/article/pii/S0890540184710376}, - author = {D. Kozen}, - abstract = {We give a finitary axiomatization of the algebra of regular events involving only - equations and equational implications. Unlike Salomaa′s axiomatizations, the - axiomatization given here is sound for all interpretations over Kleene algebras.} +@Article{ kozen1994, + title = {A Completeness Theorem for Kleene Algebras and the Algebra + of Regular Events}, + journal = {Information and Computation}, + volume = {110}, + number = {2}, + pages = {366-390}, + year = {1994}, + issn = {0890-5401}, + doi = {https://doi.org/10.1006/inco.1994.1037}, + url = {https://www.sciencedirect.com/science/article/pii/S0890540184710376}, + author = {D. Kozen}, + abstract = {We give a finitary axiomatization of the algebra of + regular events involving only equations and equational + implications. Unlike Salomaa′s axiomatizations, the + axiomatization given here is sound for all interpretations + over Kleene algebras.} } @Article{ lazarus1973, @@ -1884,7 +1888,7 @@ @Article{ salwinski2018 issn = {0746-8342}, mrclass = {26A06 (00A05)}, mrnumber = {3766700}, - doi = {10.1080/07468342.2018.1419703}, + doi = {10.1080/07468342.2018.1419703} } @Book{ samuel1967, diff --git a/scripts/lint-bib.sh b/scripts/lint-bib.sh new file mode 100755 index 0000000000000..630f64c625db2 --- /dev/null +++ b/scripts/lint-bib.sh @@ -0,0 +1,8 @@ +#!/usr/bin/env bash + +set -exo pipefail +# https://leanprover-community.github.io/contribute/doc.html#citing-other-works +cp docs/references.bib docs/references.bib.old +bibtool --preserve.key.case=on --preserve.keys=on --pass.comments=on --print.use.tab=off -s \ + -i docs/references.bib -o docs/references.bib +diff -U8 docs/references.bib.old docs/references.bib From e7286cac412124bcb9114d1403c43c8a0f644f09 Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Fri, 17 Feb 2023 06:57:46 +0000 Subject: [PATCH 0491/1291] feat(measure_theory/integral): break up integrals into sums (#18425) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This shows that if `f` is integrable on each of a countable family of measurable sets in a measure space, and the sum of the integrals of `‖f‖` over these sets converges, then `f` is integrable on their union. We also consider the specific case of breaking up `ℝ` into the union of intervals of length 1. --- .../integral/interval_integral.lean | 22 +++++++++++++++++++ src/measure_theory/integral/set_integral.lean | 22 +++++++++++++++++++ 2 files changed, 44 insertions(+) diff --git a/src/measure_theory/integral/interval_integral.lean b/src/measure_theory/integral/interval_integral.lean index baece1a595a8e..2747eebf1519c 100644 --- a/src/measure_theory/integral/interval_integral.lean +++ b/src/measure_theory/integral/interval_integral.lean @@ -1425,6 +1425,28 @@ end mono end +section has_sum +variables {μ : measure ℝ} {f : ℝ → E} + +lemma _root_.measure_theory.integrable.has_sum_interval_integral (hfi : integrable f μ) (y : ℝ) : + has_sum (λ (n : ℤ), ∫ x in (y + n)..(y + n + 1), f x ∂μ) (∫ x, f x ∂μ) := +begin + simp_rw integral_of_le (le_add_of_nonneg_right zero_le_one), + rw [←integral_univ, ←Union_Ioc_add_int_cast y], + exact has_sum_integral_Union (λ i, measurable_set_Ioc) (pairwise_disjoint_Ioc_add_int_cast y) + hfi.integrable_on, +end + +lemma _root_.measure_theory.integrable.has_sum_interval_integral_comp_add_int + (hfi : integrable f) : + has_sum (λ (n : ℤ), ∫ x in 0..1, f (x + n)) (∫ x, f x) := +begin + convert hfi.has_sum_interval_integral 0 using 2, + ext1 n, + rw [integral_comp_add_right, zero_add, add_comm], +end + +end has_sum /-! ### Fundamental theorem of calculus, part 1, for any measure diff --git a/src/measure_theory/integral/set_integral.lean b/src/measure_theory/integral/set_integral.lean index 02b274b9bf2cb..0432a39540c6a 100644 --- a/src/measure_theory/integral/set_integral.lean +++ b/src/measure_theory/integral/set_integral.lean @@ -709,6 +709,28 @@ end end nonneg +section integrable_Union + +variables {μ : measure α} [normed_add_comm_group E] {f : α → E} [countable β] {s : β → set α} + +lemma integrable_on_Union_of_summable_integral_norm + (hs : ∀ (b : β), measurable_set (s b)) (hi : ∀ (b : β), integrable_on f (s b) μ) + (h : summable (λ (b : β), ∫ (a : α) in s b, ‖f a‖ ∂μ)) : + integrable_on f (Union s) μ := +begin + refine ⟨ae_strongly_measurable.Union (λ i, (hi i).1), (lintegral_Union_le _ _).trans_lt _⟩, + have B := λ (b : β), lintegral_coe_eq_integral (λ (a : α), ‖f a‖₊) (hi b).norm, + rw tsum_congr B, + have S' : summable (λ (b : β), (⟨∫ (a : α) in s b, ‖f a‖₊ ∂μ, + set_integral_nonneg (hs b) (λ a ha, nnreal.coe_nonneg _)⟩ : nnreal)), + { rw ←nnreal.summable_coe, exact h }, + have S'' := ennreal.tsum_coe_eq S'.has_sum, + simp_rw [ennreal.coe_nnreal_eq, nnreal.coe_mk, coe_nnnorm] at S'', + convert ennreal.of_real_lt_top, +end + +end integrable_Union + section tendsto_mono variables {μ : measure α} [normed_add_comm_group E] [complete_space E] [normed_space ℝ E] From 0324bd3f61019c744cf2e88fd9c7373eb2fdb724 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 17 Feb 2023 15:56:38 +0000 Subject: [PATCH 0492/1291] fix(ring_theory/derivation): speed up a slow definition (#18457) 16s to 3s. I don't know which of the changes is responsible for the speedup, but they all seem sensible anyway. With the explicit `to_fun` field, we may as well use `simps` to generate a lemma for us. --- src/ring_theory/derivation.lean | 17 ++++++++--------- 1 file changed, 8 insertions(+), 9 deletions(-) diff --git a/src/ring_theory/derivation.lean b/src/ring_theory/derivation.lean index 940fe5076a30f..7860637d85bfe 100644 --- a/src/ring_theory/derivation.lean +++ b/src/ring_theory/derivation.lean @@ -451,29 +451,28 @@ lemma derivation_to_square_zero_of_lift_apply (f : A →ₐ[R] B) /-- Given a tower of algebras `R → A → B`, and a square-zero `I : ideal B`, each `R`-derivation from `A` to `I` corresponds to a lift `A →ₐ[R] B` of the canonical map `A →ₐ[R] B ⧸ I`. -/ +@[simps {attrs := []}] def lift_of_derivation_to_square_zero (f : derivation R A I) : A →ₐ[R] B := -{ map_one' := show (f 1 : B) + algebra_map A B 1 = 1, - by rw [map_one, f.map_one_eq_zero, submodule.coe_zero, zero_add], +{ to_fun := λ x, f x + algebra_map A B x, + map_one' := by rw [map_one, f.map_one_eq_zero, submodule.coe_zero, zero_add], map_mul' := λ x y, begin have : (f x : B) * (f y) = 0, { rw [← ideal.mem_bot, ← hI, pow_two], convert (ideal.mul_mem_mul (f x).2 (f y).2) using 1 }, - dsimp, simp only [map_mul, f.leibniz, add_mul, mul_add, submodule.coe_add, submodule.coe_smul_of_tower, algebra.smul_def, this], ring end, - commutes' := λ r, by { dsimp, simp [← is_scalar_tower.algebra_map_apply R A B r] }, + commutes' := λ r, + by simp only [derivation.map_algebra_map, eq_self_iff_true, zero_add, submodule.coe_zero, + ←is_scalar_tower.algebra_map_apply R A B r], map_zero' := ((I.restrict_scalars R).subtype.comp f.to_linear_map + (is_scalar_tower.to_alg_hom R A B).to_linear_map).map_zero, ..((I.restrict_scalars R).subtype.comp f.to_linear_map + - (is_scalar_tower.to_alg_hom R A B).to_linear_map) } - -lemma lift_of_derivation_to_square_zero_apply (f : derivation R A I) (x : A) : - lift_of_derivation_to_square_zero I hI f x = f x + algebra_map A B x := rfl + (is_scalar_tower.to_alg_hom R A B).to_linear_map : A →ₗ[R] B) } @[simp] lemma lift_of_derivation_to_square_zero_mk_apply (d : derivation R A I) (x : A) : - ideal.quotient.mk I (lift_of_derivation_to_square_zero I hI d x) = algebra_map A (B ⧸ I) x := + ideal.quotient.mk I (lift_of_derivation_to_square_zero I hI d x) = algebra_map A (B ⧸ I) x := by { rw [lift_of_derivation_to_square_zero_apply, map_add, ideal.quotient.eq_zero_iff_mem.mpr (d x).prop, zero_add], refl } From 2738d2ca56cbc63be80c3bd48e9ed90ad94e947d Mon Sep 17 00:00:00 2001 From: Michael Stoll Date: Fri, 17 Feb 2023 18:50:51 +0000 Subject: [PATCH 0493/1291] doc(algebra/continued_fractions/*): TeXify docstrings (#18459) Improve readability of continued fractions in the docstrings by randering them via LaTeX. --- src/algebra/continued_fractions/basic.lean | 54 ++++++++----------- .../approximation_corollaries.lean | 2 +- .../convergents_equiv.lean | 18 +++---- 3 files changed, 29 insertions(+), 45 deletions(-) diff --git a/src/algebra/continued_fractions/basic.lean b/src/algebra/continued_fractions/basic.lean index 6d2c43d0f36d7..643a4d8a3a99c 100644 --- a/src/algebra/continued_fractions/basic.lean +++ b/src/algebra/continued_fractions/basic.lean @@ -80,17 +80,13 @@ variable (α) /-- A *generalised continued fraction* (gcf) is a potentially infinite expression of the form - - a₀ - h + --------------------------- - a₁ - b₀ + -------------------- - a₂ - b₁ + -------------- - a₃ - b₂ + -------- - b₃ + ... - +$$ + h + \dfrac{a_0} + {b_0 + \dfrac{a_1} + {b_1 + \dfrac{a_2} + {b_2 + \dfrac{a_3} + {b_3 + \dots}}}} +$$ where `h` is called the *head term* or *integer part*, the `aᵢ` are called the *partial numerators* and the `bᵢ` the *partial denominators* of the gcf. We store the sequence of partial numerators and denominators in a sequence of @@ -150,17 +146,13 @@ end generalized_continued_fraction /-- A generalized continued fraction is a *simple continued fraction* if all partial numerators are equal to one. - - 1 - h + --------------------------- - 1 - b₀ + -------------------- - 1 - b₁ + -------------- - 1 - b₂ + -------- - b₃ + ... - +$$ + h + \dfrac{1} + {b_0 + \dfrac{1} + {b_1 + \dfrac{1} + {b_2 + \dfrac{1} + {b_3 + \dots}}}} +$$ -/ def generalized_continued_fraction.is_simple_continued_fraction (g : generalized_continued_fraction α) [has_one α] : Prop := @@ -170,17 +162,13 @@ variable (α) /-- A *simple continued fraction* (scf) is a generalized continued fraction (gcf) whose partial numerators are equal to one. - - 1 - h + --------------------------- - 1 - b₀ + -------------------- - 1 - b₁ + -------------- - 1 - b₂ + -------- - b₃ + ... - +$$ + h + \dfrac{1} + {b_0 + \dfrac{1} + {b_1 + \dfrac{1} + {b_2 + \dfrac{1} + {b_3 + \dots}}}} +$$ For convenience, one often writes `[h; b₀, b₁, b₂,...]`. It is encoded as the subtype of gcfs that satisfy `generalized_continued_fraction.is_simple_continued_fraction`. diff --git a/src/algebra/continued_fractions/computation/approximation_corollaries.lean b/src/algebra/continued_fractions/computation/approximation_corollaries.lean index 01f15e9e1a730..9c22e4f9aa49b 100644 --- a/src/algebra/continued_fractions/computation/approximation_corollaries.lean +++ b/src/algebra/continued_fractions/computation/approximation_corollaries.lean @@ -70,7 +70,7 @@ lemma of_convergents_eq_convergents' : The recurrence relation for the `convergents` of the continued fraction expansion of an element `v` of `K` in terms of the convergents of the inverse of its fractional part. -/ -lemma convergents_succ (n : ℕ): +lemma convergents_succ (n : ℕ) : (of v).convergents (n + 1) = ⌊v⌋ + 1 / (of (int.fract v)⁻¹).convergents n := by rw [of_convergents_eq_convergents', convergents'_succ, of_convergents_eq_convergents'] diff --git a/src/algebra/continued_fractions/convergents_equiv.lean b/src/algebra/continued_fractions/convergents_equiv.lean index 439bf0d209637..244c9e843532a 100644 --- a/src/algebra/continued_fractions/convergents_equiv.lean +++ b/src/algebra/continued_fractions/convergents_equiv.lean @@ -18,17 +18,13 @@ direct evaluation (`convergents'`)) for `gcf`s on linear ordered fields. We foll [hardy2008introduction], Chapter 10. Here's a sketch: Let `c` be a continued fraction `[h; (a₀, b₀), (a₁, b₁), (a₂, b₂),...]`, visually: - - a₀ - h + --------------------------- - a₁ - b₀ + -------------------- - a₂ - b₁ + -------------- - a₃ - b₂ + -------- - b₃ + ... - +$$ + c = h + \dfrac{a_0} + {b_0 + \dfrac{a_1} + {b_1 + \dfrac{a_2} + {b_2 + \dfrac{a_3} + {b_3 + \dots}}}} +$$ One can compute the convergents of `c` in two ways: 1. Directly evaluating the fraction described by `c` up to a given `n` (`convergents'`) 2. Using the recurrence (`convergents`): From 297619ec79dedf23525458b6bf5bf35c736fd2b8 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sat, 18 Feb 2023 08:47:16 +0000 Subject: [PATCH 0494/1291] chore(number_theory/sum_four_squares): squeeze simps (#18461) This proof goes from 20s to 8s. Adding a missing `norm_cast` lemma makes it possible to replace some `simp`s with `push_cast` --- src/data/zmod/basic.lean | 2 +- src/number_theory/sum_four_squares.lean | 27 ++++++++++++++----------- 2 files changed, 16 insertions(+), 13 deletions(-) diff --git a/src/data/zmod/basic.lean b/src/data/zmod/basic.lean index 8c8957dd63ae9..603f387fe703e 100644 --- a/src/data/zmod/basic.lean +++ b/src/data/zmod/basic.lean @@ -778,7 +778,7 @@ begin { refl } end -@[simp] lemma coe_val_min_abs : ∀ {n : ℕ} (x : zmod n), (x.val_min_abs : zmod n) = x +@[simp, norm_cast] lemma coe_val_min_abs : ∀ {n : ℕ} (x : zmod n), (x.val_min_abs : zmod n) = x | 0 x := int.cast_id x | k@(n+1) x := begin diff --git a/src/number_theory/sum_four_squares.lean b/src/number_theory/sum_four_squares.lean index 39c58c6f36e3c..34800f552522d 100644 --- a/src/number_theory/sum_four_squares.lean +++ b/src/number_theory/sum_four_squares.lean @@ -28,7 +28,7 @@ namespace int lemma sq_add_sq_of_two_mul_sq_add_sq {m x y : ℤ} (h : 2 * m = x^2 + y^2) : m = ((x - y) / 2) ^ 2 + ((x + y) / 2) ^ 2 := -have even (x^2 + y^2), by simp [h.symm, even_mul], +have even (x^2 + y^2), by simp [←h, even_mul], have hxaddy : even (x + y), by simpa [sq] with parity_simps, have hxsuby : even (x - y), by simpa [sq] with parity_simps, (mul_right_inj' (show (2*2 : ℤ) ≠ 0, from dec_trivial)).1 $ @@ -113,7 +113,7 @@ have hm : ∃ m < p, 0 < m ∧ ∃ a b c d : ℤ, a^2 + b^2 + c^2 + d^2 = m * p, (λ hk0, by { rw [hk0, int.coe_nat_zero, zero_mul] at hk, exact ne_of_gt (show a^2 + b^2 + 1 > 0, from add_pos_of_nonneg_of_pos (add_nonneg (sq_nonneg _) (sq_nonneg _)) zero_lt_one) hk.1 }), - a, b, 1, 0, by simpa [sq] using hk.1⟩, + a, b, 1, 0, by simpa only [zero_pow two_pos, one_pow, add_zero] using hk.1⟩, let m := nat.find hm in let ⟨a, b, c, d, (habcd : a^2 + b^2 + c^2 + d^2 = m * p)⟩ := (nat.find_spec hm).snd.2 in by haveI hm0 : ne_zero m := ne_zero.of_pos (nat.find_spec hm).snd.1; exact @@ -121,7 +121,8 @@ have hmp : m < p, from (nat.find_spec hm).fst, m.mod_two_eq_zero_or_one.elim (λ hm2 : m % 2 = 0, let ⟨k, hk⟩ := nat.dvd_iff_mod_eq_zero.2 hm2 in - have hk0 : 0 < k, from nat.pos_of_ne_zero $ λ _, by { simp [*, lt_irrefl] at * }, + have hk0 : 0 < k, from nat.pos_of_ne_zero $ + by { rintro rfl, rw mul_zero at hk, exact ne_zero.ne m hk }, have hkm : k < m, { rw [hk, two_mul], exact (lt_add_iff_pos_left _).2 hk0 }, false.elim $ nat.find_min hm hkm ⟨lt_trans hkm hmp, hk0, sum_four_squares_of_two_mul_sum_four_squares @@ -134,7 +135,7 @@ m.mod_two_eq_zero_or_one.elim y := (c : zmod m).val_min_abs, z := (d : zmod m).val_min_abs in have hnat_abs : w^2 + x^2 + y^2 + z^2 = (w.nat_abs^2 + x.nat_abs^2 + y.nat_abs ^2 + z.nat_abs ^ 2 : ℕ), - by simp [sq], + by { push_cast, simp_rw sq_abs, }, have hwxyzlt : w^2 + x^2 + y^2 + z^2 < m^2, from calc w^2 + x^2 + y^2 + z^2 = (w.nat_abs^2 + x.nat_abs^2 + y.nat_abs ^2 + z.nat_abs ^ 2 : ℕ) : hnat_abs @@ -144,7 +145,8 @@ m.mod_two_eq_zero_or_one.elim (nat.pow_le_pow_of_le_left (zmod.nat_abs_val_min_abs_le _) _)) (nat.pow_le_pow_of_le_left (zmod.nat_abs_val_min_abs_le _) _)) (nat.pow_le_pow_of_le_left (zmod.nat_abs_val_min_abs_le _) _) - ... = 4 * (m / 2 : ℕ) ^ 2 : by simp [sq, bit0, bit1, mul_add, add_mul, add_assoc] + ... = 4 * (m / 2 : ℕ) ^ 2 : by simp only [bit0_mul, one_mul, two_smul, + nat.cast_add, nat.cast_pow, add_assoc] ... < 4 * (m / 2 : ℕ) ^ 2 + ((4 * (m / 2) : ℕ) * (m % 2 : ℕ) + (m % 2 : ℕ)^2) : (lt_add_iff_pos_right _).2 (by { rw [hm2, int.coe_nat_one, one_pow, mul_one], exact add_pos_of_nonneg_of_pos (int.coe_nat_nonneg _) zero_lt_one }) @@ -153,7 +155,7 @@ m.mod_two_eq_zero_or_one.elim pow_add, add_comm, add_left_comm] }, have hwxyzabcd : ((w^2 + x^2 + y^2 + z^2 : ℤ) : zmod m) = ((a^2 + b^2 + c^2 + d^2 : ℤ) : zmod m), - by simp [w, x, y, z, sq], + by push_cast, have hwxyz0 : ((w^2 + x^2 + y^2 + z^2 : ℤ) : zmod m) = 0, by rw [hwxyzabcd, habcd, int.cast_mul, cast_coe_nat, zmod.nat_cast_self, zero_mul], let ⟨n, hn⟩ := ((char_p.int_cast_eq_zero_iff _ m _).1 hwxyz0) in @@ -161,8 +163,8 @@ m.mod_two_eq_zero_or_one.elim have hwxyz0 : (w.nat_abs^2 + x.nat_abs^2 + y.nat_abs^2 + z.nat_abs^2 : ℕ) = 0, by { rw [← int.coe_nat_eq_zero, ← hnat_abs], rwa [hn0, mul_zero] at hn }, have habcd0 : (m : ℤ) ∣ a ∧ (m : ℤ) ∣ b ∧ (m : ℤ) ∣ c ∧ (m : ℤ) ∣ d, - by simpa [add_eq_zero_iff' (sq_nonneg (_ : ℤ)) (sq_nonneg _), - pow_two, w, x, y, z, (char_p.int_cast_eq_zero_iff _ m _), and.assoc] using hwxyz0, + by simpa only [add_eq_zero_iff, int.nat_abs_eq_zero, zmod.val_min_abs_eq_zero, and.assoc, + pow_eq_zero_iff two_pos, char_p.int_cast_eq_zero_iff _ m _] using hwxyz0, let ⟨ma, hma⟩ := habcd0.1, ⟨mb, hmb⟩ := habcd0.2.1, ⟨mc, hmc⟩ := habcd0.2.2.1, ⟨md, hmd⟩ := habcd0.2.2.2 in have hmdvdp : m ∣ p, @@ -172,13 +174,14 @@ m.mod_two_eq_zero_or_one.elim (hp.1.eq_one_or_self_of_dvd _ hmdvdp).elim hm1 (λ hmeqp, by simpa [lt_irrefl, hmeqp] using hmp)), have hawbxcydz : ((m : ℕ) : ℤ) ∣ a * w + b * x + c * y + d * z, - from (char_p.int_cast_eq_zero_iff (zmod m) m _).1 $ by { rw [← hwxyz0], simp, ring }, + from (char_p.int_cast_eq_zero_iff (zmod m) m _).1 $ + by { rw [← hwxyz0], simp_rw [sq], push_cast }, have haxbwczdy : ((m : ℕ) : ℤ) ∣ a * x - b * w - c * z + d * y, - from (char_p.int_cast_eq_zero_iff (zmod m) m _).1 $ by { simp [sub_eq_add_neg], ring }, + from (char_p.int_cast_eq_zero_iff (zmod m) m _).1 $ by { push_cast, ring }, have haybzcwdx : ((m : ℕ) : ℤ) ∣ a * y + b * z - c * w - d * x, - from (char_p.int_cast_eq_zero_iff (zmod m) m _).1 $ by { simp [sub_eq_add_neg], ring }, + from (char_p.int_cast_eq_zero_iff (zmod m) m _).1 $ by { push_cast, ring }, have hazbycxdw : ((m : ℕ) : ℤ) ∣ a * z - b * y + c * x - d * w, - from (char_p.int_cast_eq_zero_iff (zmod m) m _).1 $ by { simp [sub_eq_add_neg], ring }, + from (char_p.int_cast_eq_zero_iff (zmod m) m _).1 $ by { push_cast, ring }, let ⟨s, hs⟩ := hawbxcydz, ⟨t, ht⟩ := haxbwczdy, ⟨u, hu⟩ := haybzcwdx, ⟨v, hv⟩ := hazbycxdw in have hn_nonneg : 0 ≤ n, from nonneg_of_mul_nonneg_right From e97cf15cd1aec9bd5c193b2ffac5a6dc9118912b Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Sat, 18 Feb 2023 19:08:45 +0000 Subject: [PATCH 0495/1291] chore(*): add mathlib4 synchronization comments (#18450) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `category_theory.arrow` * `category_theory.balanced` * `category_theory.category.preorder` * `category_theory.endomorphism` * `category_theory.epi_mono` * `category_theory.functor.currying` * `category_theory.functor.hom` * `category_theory.groupoid` * `category_theory.groupoid.basic` * `category_theory.pi.basic` * `category_theory.types` * `computability.DFA` * `computability.language` * `computability.turing_machine` * `data.fintype.card_embedding` * `group_theory.commensurable` * `order.filter.germ` * `set_theory.ordinal.arithmetic` * `topology.algebra.group_with_zero` * `topology.list` --- src/category_theory/arrow.lean | 3 +++ src/category_theory/balanced.lean | 3 +++ src/category_theory/category/preorder.lean | 3 +++ src/category_theory/endomorphism.lean | 3 +++ src/category_theory/epi_mono.lean | 3 +++ src/category_theory/functor/currying.lean | 3 +++ src/category_theory/functor/hom.lean | 3 +++ src/category_theory/groupoid.lean | 3 +++ src/category_theory/groupoid/basic.lean | 3 +++ src/category_theory/pi/basic.lean | 3 +++ src/category_theory/types.lean | 3 +++ src/computability/DFA.lean | 3 +++ src/computability/language.lean | 3 +++ src/computability/turing_machine.lean | 3 +++ src/data/fintype/card_embedding.lean | 3 +++ src/group_theory/commensurable.lean | 3 +++ src/order/filter/germ.lean | 3 +++ src/set_theory/ordinal/arithmetic.lean | 3 +++ src/topology/algebra/group_with_zero.lean | 3 +++ src/topology/list.lean | 3 +++ 20 files changed, 60 insertions(+) diff --git a/src/category_theory/arrow.lean b/src/category_theory/arrow.lean index 51dfed31145cc..3df992cb8a5d0 100644 --- a/src/category_theory/arrow.lean +++ b/src/category_theory/arrow.lean @@ -8,6 +8,9 @@ import category_theory.comma /-! # The category of arrows +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The category of arrows, with morphisms commutative squares. We set this up as a specialization of the comma category `comma L R`, where `L` and `R` are both the identity functor. diff --git a/src/category_theory/balanced.lean b/src/category_theory/balanced.lean index c1277508f7cef..2053a3ef7ff26 100644 --- a/src/category_theory/balanced.lean +++ b/src/category_theory/balanced.lean @@ -8,6 +8,9 @@ import category_theory.epi_mono /-! # Balanced categories +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A category is called balanced if any morphism that is both monic and epic is an isomorphism. Balanced categories arise frequently. For example, categories in which every monomorphism diff --git a/src/category_theory/category/preorder.lean b/src/category_theory/category/preorder.lean index cac31c7f2eea2..46f3bcc89b813 100644 --- a/src/category_theory/category/preorder.lean +++ b/src/category_theory/category/preorder.lean @@ -10,6 +10,9 @@ import order.hom.basic # Preorders as categories +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We install a category instance on any preorder. This is not to be confused with the category _of_ preorders, defined in `order/category/Preorder`. diff --git a/src/category_theory/endomorphism.lean b/src/category_theory/endomorphism.lean index d65377635225f..47c1aa2b5eff0 100644 --- a/src/category_theory/endomorphism.lean +++ b/src/category_theory/endomorphism.lean @@ -11,6 +11,9 @@ import group_theory.group_action.defs /-! # Endomorphisms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Definition and basic properties of endomorphisms and automorphisms of an object in a category. For each `X : C`, we provide `End X := X ⟶ X` with a monoid structure, diff --git a/src/category_theory/epi_mono.lean b/src/category_theory/epi_mono.lean index ee62732202535..b23311392e635 100644 --- a/src/category_theory/epi_mono.lean +++ b/src/category_theory/epi_mono.lean @@ -9,6 +9,9 @@ import category_theory.groupoid /-! # Facts about epimorphisms and monomorphisms. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The definitions of `epi` and `mono` are in `category_theory.category`, since they are used by some lemmas for `iso`, which is used everywhere. -/ diff --git a/src/category_theory/functor/currying.lean b/src/category_theory/functor/currying.lean index 26a95ef824fa7..a0c3fa5094f04 100644 --- a/src/category_theory/functor/currying.lean +++ b/src/category_theory/functor/currying.lean @@ -8,6 +8,9 @@ import category_theory.products.bifunctor /-! # Curry and uncurry, as functors. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define `curry : ((C × D) ⥤ E) ⥤ (C ⥤ (D ⥤ E))` and `uncurry : (C ⥤ (D ⥤ E)) ⥤ ((C × D) ⥤ E)`, and verify that they provide an equivalence of categories `currying : (C ⥤ (D ⥤ E)) ≌ ((C × D) ⥤ E)`. diff --git a/src/category_theory/functor/hom.lean b/src/category_theory/functor/hom.lean index ee307cb400b05..e5f7512a747a5 100644 --- a/src/category_theory/functor/hom.lean +++ b/src/category_theory/functor/hom.lean @@ -7,6 +7,9 @@ import category_theory.products.basic import category_theory.types /-! +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The hom functor, sending `(X, Y)` to the type `X ⟶ Y`. -/ diff --git a/src/category_theory/groupoid.lean b/src/category_theory/groupoid.lean index c3b0ca0a31ab6..8f6b2105beea3 100644 --- a/src/category_theory/groupoid.lean +++ b/src/category_theory/groupoid.lean @@ -12,6 +12,9 @@ import combinatorics.quiver.connected_component /-! # Groupoids +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define `groupoid` as a typeclass extending `category`, asserting that all morphisms have inverses. diff --git a/src/category_theory/groupoid/basic.lean b/src/category_theory/groupoid/basic.lean index 1cf1a3bb07e88..edcd548ec99a5 100644 --- a/src/category_theory/groupoid/basic.lean +++ b/src/category_theory/groupoid/basic.lean @@ -7,6 +7,9 @@ import category_theory.groupoid import combinatorics.quiver.basic /-! +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines a few basic properties of groupoids. -/ diff --git a/src/category_theory/pi/basic.lean b/src/category_theory/pi/basic.lean index 5389104a9d34e..f5f844e1027c9 100644 --- a/src/category_theory/pi/basic.lean +++ b/src/category_theory/pi/basic.lean @@ -10,6 +10,9 @@ import data.sum.basic /-! # Categories of indexed families of objects. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define the pointwise category structure on indexed families of objects in a category (and also the dependent generalization). diff --git a/src/category_theory/types.lean b/src/category_theory/types.lean index 1f16fcbe2ad27..ad97ce8539f57 100644 --- a/src/category_theory/types.lean +++ b/src/category_theory/types.lean @@ -10,6 +10,9 @@ import logic.equiv.basic /-! # The category `Type`. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this section we set up the theory so that Lean's types and functions between them can be viewed as a `large_category` in our framework. diff --git a/src/computability/DFA.lean b/src/computability/DFA.lean index 948b28ad9ce6c..2bf2fadcda8a2 100644 --- a/src/computability/DFA.lean +++ b/src/computability/DFA.lean @@ -11,6 +11,9 @@ import tactic.norm_num /-! # Deterministic Finite Automata +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains the definition of a Deterministic Finite Automaton (DFA), a state machine which determines whether a string (implemented as a list over an arbitrary alphabet) is in a regular set in linear time. diff --git a/src/computability/language.lean b/src/computability/language.lean index ae2ef1a0e5ef4..c1c741e9fc420 100644 --- a/src/computability/language.lean +++ b/src/computability/language.lean @@ -11,6 +11,9 @@ import data.set.lattice /-! # Languages +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains the definition and operations on formal languages over an alphabet. Note strings are implemented as lists over the alphabet. The operations in this file define a [Kleene algebra](https://en.wikipedia.org/wiki/Kleene_algebra) diff --git a/src/computability/turing_machine.lean b/src/computability/turing_machine.lean index 4bb8e26ed37fc..e8a283b317b00 100644 --- a/src/computability/turing_machine.lean +++ b/src/computability/turing_machine.lean @@ -15,6 +15,9 @@ import tactic.apply_fun /-! # Turing machines +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines a sequence of simple machine languages, starting with Turing machines and working up to more complex languages based on Wang B-machines. diff --git a/src/data/fintype/card_embedding.lean b/src/data/fintype/card_embedding.lean index 9992b97aef63c..e4b885e127971 100644 --- a/src/data/fintype/card_embedding.lean +++ b/src/data/fintype/card_embedding.lean @@ -9,6 +9,9 @@ import logic.equiv.embedding /-! # Number of embeddings +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file establishes the cardinality of `α ↪ β` in full generality. -/ diff --git a/src/group_theory/commensurable.lean b/src/group_theory/commensurable.lean index 1d7772b62ef73..bf7f219b618b5 100644 --- a/src/group_theory/commensurable.lean +++ b/src/group_theory/commensurable.lean @@ -9,6 +9,9 @@ import group_theory.group_action.conj_act /-! # Commensurability for subgroups +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines commensurability for subgroups of a group `G`. It then goes on to prove that commensurability defines an equivalence relation and finally defines the commensurator of a subgroup of `G`. diff --git a/src/order/filter/germ.lean b/src/order/filter/germ.lean index 2a6c20199d729..c6f8199fe2172 100644 --- a/src/order/filter/germ.lean +++ b/src/order/filter/germ.lean @@ -9,6 +9,9 @@ import algebra.module.pi /-! # Germ of a function at a filter +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The germ of a function `f : α → β` at a filter `l : filter α` is the equivalence class of `f` with respect to the equivalence relation `eventually_eq l`: `f ≈ g` means `∀ᶠ x in l, f x = g x`. diff --git a/src/set_theory/ordinal/arithmetic.lean b/src/set_theory/ordinal/arithmetic.lean index 7a3bd7d4615cf..198622e437622 100644 --- a/src/set_theory/ordinal/arithmetic.lean +++ b/src/set_theory/ordinal/arithmetic.lean @@ -9,6 +9,9 @@ import tactic.by_contra /-! # Ordinal arithmetic +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Ordinals have an addition (corresponding to disjoint union) that turns them into an additive monoid, and a multiplication (corresponding to the lexicographic order on the product) that turns them into a monoid. One can also define correspondingly a subtraction, a division, a successor diff --git a/src/topology/algebra/group_with_zero.lean b/src/topology/algebra/group_with_zero.lean index f8f1f47ac7c9b..15b818a7865a7 100644 --- a/src/topology/algebra/group_with_zero.lean +++ b/src/topology/algebra/group_with_zero.lean @@ -10,6 +10,9 @@ import topology.homeomorph /-! # Topological group with zero +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define `has_continuous_inv₀` to be a mixin typeclass a type with `has_inv` and `has_zero` (e.g., a `group_with_zero`) such that `λ x, x⁻¹` is continuous at all nonzero points. Any normed (semi)field has this property. Currently the only example of `has_continuous_inv₀` in diff --git a/src/topology/list.lean b/src/topology/list.lean index 0626f4cdf25c7..9c9c61c317aa0 100644 --- a/src/topology/list.lean +++ b/src/topology/list.lean @@ -8,6 +8,9 @@ import topology.algebra.monoid /-! # Topology on lists and vectors +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + -/ open topological_space set filter open_locale topology filter From 28aa996fc6fb4317f0083c4e6daf79878d81be33 Mon Sep 17 00:00:00 2001 From: Jim Fowler Date: Sun, 19 Feb 2023 17:32:40 +0000 Subject: [PATCH 0496/1291] feat(algebra/quandle): Unital shelves (#17839) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Extend shelves to include the unital case, as well as a few basic results about unital shelves from https://arxiv.org/pdf/1603.08590.pdf This culminates in providing a monoid instance for a unital shelf. Co-authored-by: Jim Fowler Co-authored-by: Yaël Dillies --- docs/references.bib | 18 +++++++++++++++ src/algebra/quandle.lean | 48 +++++++++++++++++++++++++++++++++++++++- 2 files changed, 65 insertions(+), 1 deletion(-) diff --git a/docs/references.bib b/docs/references.bib index acbfb060539cc..9bd399e6dfc22 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -600,6 +600,24 @@ @Book{ coxlittleoshea1997 isbn = {978-0-387-94680-1} } +@Article{ crans2017, + author = {Crans, Alissa S. and Mukherjee, Sujoy and Przytycki, + J\'{o}zef H.}, + title = {On homology of associative shelves}, + journal = {J. Homotopy Relat. Struct.}, + fjournal = {Journal of Homotopy and Related Structures}, + volume = {12}, + year = {2017}, + number = {3}, + pages = {741--763}, + issn = {2193-8407}, + mrclass = {18G60 (20M32 20N02 57M25)}, + mrnumber = {3691304}, + mrreviewer = {Mahender Singh}, + doi = {10.1007/s40062-016-0164-9}, + url = {https://doi.org/10.1007/s40062-016-0164-9} +} + @Book{ davey_priestley, author = {Davey, B. A. and Priestley, H. A.}, title = {Introduction to lattices and order}, diff --git a/src/algebra/quandle.lean b/src/algebra/quandle.lean index a1e9b87f20d72..6e98bb487b0e9 100644 --- a/src/algebra/quandle.lean +++ b/src/algebra/quandle.lean @@ -34,7 +34,7 @@ complements that is analogous to the fundamental group of the exterior, and he showed that the quandle associated to an oriented knot is invariant up to orientation-reversed mirror image. Racks were used by Fenn and Rourke for framed codimension-2 knots and -links.[FennRourke1992] +links in [FennRourke1992]. Unital shelves are discussed in [crans2017]. The name "rack" came from wordplay by Conway and Wraith for the "wrack and ruin" of forgetting everything but the conjugation operation for a @@ -43,6 +43,7 @@ group. ## Main definitions * `shelf` is a type with a self-distributive action +* `unital_shelf` is a shelf with a left and right unit * `rack` is a shelf whose action for each element is invertible * `quandle` is a rack whose action for an element fixes that element * `quandle.conj` defines a quandle of a group acting on itself by conjugation. @@ -54,6 +55,11 @@ group. * `rack.envel_group` is left adjoint to `quandle.conj` (`to_envel_group.map`). The universality statements are `to_envel_group.univ` and `to_envel_group.univ_uniq`. +## Implementation notes + +"Unital racks" are uninteresting (see `rack.assoc_iff_id`, `unital_shelf.assoc`), so we do not +define them. + ## Notation The following notation is localized in `quandles`: @@ -92,6 +98,14 @@ class shelf (α : Type u) := (act : α → α → α) (self_distrib : ∀ {x y z : α}, act x (act y z) = act (act x y) (act x z)) +/-- +A *unital shelf* is a shelf equipped with an element `1` such that, for all elements `x`, +we have both `x ◃ 1` and `1 ◃ x` equal `x`. +-/ +class unital_shelf (α : Type u) extends shelf α, has_one α := +(one_act : ∀ a : α, act 1 a = a) +(act_one : ∀ a : α, act a 1 = a) + /-- The type of homomorphisms between shelves. This is also the notion of rack and quandle homomorphisms. @@ -120,6 +134,38 @@ localized "infixr (name := shelf_hom) ` →◃ `:25 := shelf_hom" in quandles open_locale quandles +namespace unital_shelf +open shelf + +variables {S : Type*} [unital_shelf S] + +/-- +A monoid is *graphic* if, for all `x` and `y`, the *graphic identity* +`(x * y) * x = x * y` holds. For a unital shelf, this graphic +identity holds. +-/ +lemma act_act_self_eq (x y : S) : (x ◃ y) ◃ x = x ◃ y := +begin + have h : (x ◃ y) ◃ x = (x ◃ y) ◃ (x ◃ 1) := by rw act_one, + rw [h, ←shelf.self_distrib, act_one], +end + +lemma act_idem (x : S) : (x ◃ x) = x := by rw [←act_one x, ←shelf.self_distrib, act_one, act_one] + +lemma act_self_act_eq (x y : S) : x ◃ (x ◃ y) = x ◃ y := +begin + have h : x ◃ (x ◃ y) = (x ◃ 1) ◃ (x ◃ y) := by rw act_one, + rw [h, ←shelf.self_distrib, one_act], +end + +/-- +The associativity of a unital shelf comes for free. +-/ +lemma assoc (x y z : S) : (x ◃ y) ◃ z = x ◃ y ◃ z := +by rw [self_distrib, self_distrib, act_act_self_eq, act_self_act_eq] + +end unital_shelf + namespace rack variables {R : Type*} [rack R] From 31a8a27684ce9a5749914f4248c3f7bf76605d41 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Mon, 20 Feb 2023 03:03:44 +0000 Subject: [PATCH 0497/1291] feat(probability/kernel/basic): add kernel.with_density (#18463) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Kernel with image `(κ a).with_density (f a)` if `function.uncurry f` is measurable, and with image 0 otherwise. If `function.uncurry f` is measurable, it satisfies `∫⁻ b, g b ∂(with_density κ f hf a) = ∫⁻ b, f a b * g b ∂(κ a)`. Co-authored-by: Rémy Degenne --- src/measure_theory/integral/lebesgue.lean | 14 ++ src/probability/kernel/basic.lean | 226 +++++++++++++++++++++- 2 files changed, 237 insertions(+), 3 deletions(-) diff --git a/src/measure_theory/integral/lebesgue.lean b/src/measure_theory/integral/lebesgue.lean index f90e0922cd070..8d2c15e95793a 100644 --- a/src/measure_theory/integral/lebesgue.lean +++ b/src/measure_theory/integral/lebesgue.lean @@ -2415,6 +2415,20 @@ lemma with_density_add_right (f : α → ℝ≥0∞) {g : α → ℝ≥0∞} (hg μ.with_density (f + g) = μ.with_density f + μ.with_density g := by simpa only [add_comm] using with_density_add_left hg f +lemma with_density_add_measure {m : measurable_space α} (μ ν : measure α) (f : α → ℝ≥0∞) : + (μ + ν).with_density f = μ.with_density f + ν.with_density f := +begin + ext1 s hs, + simp only [with_density_apply f hs, restrict_add, lintegral_add_measure, measure.add_apply], +end + +lemma with_density_sum {ι : Type*} {m : measurable_space α} (μ : ι → measure α) (f : α → ℝ≥0∞) : + (sum μ).with_density f = sum (λ n, (μ n).with_density f) := +begin + ext1 s hs, + simp_rw [sum_apply _ hs, with_density_apply f hs, restrict_sum μ hs, lintegral_sum_measure], +end + lemma with_density_smul (r : ℝ≥0∞) {f : α → ℝ≥0∞} (hf : measurable f) : μ.with_density (r • f) = r • μ.with_density f := begin diff --git a/src/probability/kernel/basic.lean b/src/probability/kernel/basic.lean index 6dba6cfff2e73..6a51e139debc6 100644 --- a/src/probability/kernel/basic.lean +++ b/src/probability/kernel/basic.lean @@ -27,6 +27,18 @@ Classes of kernels: stronger condition is necessary to ensure that the composition of two finite kernels is finite. * `is_s_finite_kernel κ`: a kernel is called s-finite if it is a countable sum of finite kernels. +Particular kernels: +* `deterministic {f : α → β} (hf : measurable f)`: kernel `a ↦ measure.dirac (f a)`. +* `const α (μβ : measure β)`: constant kernel `a ↦ μβ`. +* `kernel.restrict κ (hs : measurable_set s)`: kernel for which the image of `a : α` is + `(κ a).restrict s`. + Integral: `∫⁻ b, f b ∂(kernel.restrict κ hs a) = ∫⁻ b in s, f b ∂(κ a)` +* `kernel.with_density κ (f : α → β → ℝ≥0∞)`: kernel `a ↦ (κ a).with_density (f a)`. + It is defined if `κ` is s-finite. If `f` is finite everywhere, then this is also an s-finite + kernel. The class of s-finite kernels is the smallest class of kernels that contains finite + kernels and which is stable by `with_density`. + Integral: `∫⁻ b, g b ∂(with_density κ f a) = ∫⁻ b, f a b * g b ∂(κ a)` + ## Main statements * `ext_fun`: if `∫⁻ b, f b ∂(κ a) = ∫⁻ b, f b ∂(η a)` for all measurable functions `f` and all `a`, @@ -40,7 +52,7 @@ Classes of kernels: open measure_theory -open_locale measure_theory ennreal big_operators +open_locale measure_theory ennreal nnreal big_operators namespace probability_theory @@ -186,6 +198,14 @@ lemma sum_apply' [countable ι] (κ : ι → kernel α β) (a : α) {s : set β} kernel.sum κ a s = ∑' n, κ n a s := by rw [sum_apply κ a, measure.sum_apply _ hs] +@[simp] +lemma sum_zero [countable ι] : kernel.sum (λ (i : ι), (0 : kernel α β)) = 0 := +begin + ext a s hs : 2, + rw [sum_apply' _ a hs], + simp only [zero_apply, measure.coe_zero, pi.zero_apply, tsum_zero], +end + lemma sum_comm [countable ι] (κ : ι → ι → kernel α β) : kernel.sum (λ n, kernel.sum (κ n)) = kernel.sum (λ m, kernel.sum (λ n, κ n m)) := by { ext a s hs, simp_rw [sum_apply], rw measure.sum_comm, } @@ -304,8 +324,7 @@ def deterministic {f : α → β} (hf : measurable f) : begin refine measure.measurable_of_measurable_coe _ (λ s hs, _), simp_rw measure.dirac_apply' _ hs, - refine measurable.indicator _ (hf hs), - simp only [pi.one_apply, measurable_const], + exact measurable_one.indicator (hf hs), end, } lemma deterministic_apply {f : α → β} (hf : measurable f) (a : α) : @@ -536,6 +555,207 @@ measurable_set_lintegral κ (hf.comp measurable_snd) hs end measurable_lintegral +section with_density +variables {f : α → β → ℝ≥0∞} + +/-- Kernel with image `(κ a).with_density (f a)` if `function.uncurry f` is measurable, and +with image 0 otherwise. If `function.uncurry f` is measurable, it satisfies +`∫⁻ b, g b ∂(with_density κ f hf a) = ∫⁻ b, f a b * g b ∂(κ a)`. -/ +noncomputable +def with_density (κ : kernel α β) [is_s_finite_kernel κ] (f : α → β → ℝ≥0∞) : + kernel α β := +@dite _ (measurable (function.uncurry f)) (classical.dec _) + (λ hf, ({ val := λ a, (κ a).with_density (f a), + property := + begin + refine measure.measurable_of_measurable_coe _ (λ s hs, _), + simp_rw with_density_apply _ hs, + exact measurable_set_lintegral κ hf hs, + end, } : kernel α β)) + (λ hf, 0) + +lemma with_density_of_not_measurable (κ : kernel α β) [is_s_finite_kernel κ] + (hf : ¬ measurable (function.uncurry f)) : + with_density κ f = 0 := +by { classical, exact dif_neg hf, } + +protected lemma with_density_apply (κ : kernel α β) [is_s_finite_kernel κ] + (hf : measurable (function.uncurry f)) (a : α) : + with_density κ f a = (κ a).with_density (f a) := +by { classical, rw [with_density, dif_pos hf], refl, } + +lemma with_density_apply' (κ : kernel α β) [is_s_finite_kernel κ] + (hf : measurable (function.uncurry f)) (a : α) {s : set β} (hs : measurable_set s) : + with_density κ f a s = ∫⁻ b in s, f a b ∂(κ a) := +by rw [kernel.with_density_apply κ hf, with_density_apply _ hs] + +lemma lintegral_with_density (κ : kernel α β) [is_s_finite_kernel κ] + (hf : measurable (function.uncurry f)) (a : α) {g : β → ℝ≥0∞} (hg : measurable g) : + ∫⁻ b, g b ∂(with_density κ f a) = ∫⁻ b, f a b * g b ∂(κ a) := +begin + rw [kernel.with_density_apply _ hf, + lintegral_with_density_eq_lintegral_mul _ (measurable.of_uncurry_left hf) hg], + simp_rw pi.mul_apply, +end + +lemma with_density_add_left (κ η : kernel α β) [is_s_finite_kernel κ] [is_s_finite_kernel η] + (f : α → β → ℝ≥0∞) : + with_density (κ + η) f = with_density κ f + with_density η f := +begin + by_cases hf : measurable (function.uncurry f), + { ext a s hs : 2, + simp only [kernel.with_density_apply _ hf, coe_fn_add, pi.add_apply, with_density_add_measure, + measure.add_apply], }, + { simp_rw [with_density_of_not_measurable _ hf], + rw zero_add, }, +end + +lemma with_density_kernel_sum [countable ι] (κ : ι → kernel α β) + (hκ : ∀ i, is_s_finite_kernel (κ i)) (f : α → β → ℝ≥0∞) : + @with_density _ _ _ _ (kernel.sum κ) (is_s_finite_kernel_sum hκ) f + = kernel.sum (λ i, with_density (κ i) f) := +begin + by_cases hf : measurable (function.uncurry f), + { ext1 a, + simp_rw [sum_apply, kernel.with_density_apply _ hf, sum_apply, + with_density_sum (λ n, κ n a) (f a)], }, + { simp_rw [with_density_of_not_measurable _ hf], + exact sum_zero.symm, }, +end + +lemma with_density_tsum [countable ι] (κ : kernel α β) [is_s_finite_kernel κ] + {f : ι → α → β → ℝ≥0∞} (hf : ∀ i, measurable (function.uncurry (f i))) : + with_density κ (∑' n, f n) = kernel.sum (λ n, with_density κ (f n)) := +begin + have h_sum_a : ∀ a, summable (λ n, f n a) := λ a, pi.summable.mpr (λ b, ennreal.summable), + have h_sum : summable (λ n, f n) := pi.summable.mpr h_sum_a, + ext a s hs : 2, + rw [sum_apply' _ a hs, with_density_apply' κ _ a hs], + swap, + { have : function.uncurry (∑' n, f n) = ∑' n, function.uncurry (f n), + { ext1 p, + simp only [function.uncurry_def], + rw [tsum_apply h_sum, tsum_apply (h_sum_a _), tsum_apply], + exact pi.summable.mpr (λ p, ennreal.summable), }, + rw this, + exact measurable.ennreal_tsum' hf, }, + have : ∫⁻ b in s, (∑' n, f n) a b ∂(κ a) = ∫⁻ b in s, (∑' n, (λ b, f n a b) b) ∂(κ a), + { congr' with b, + rw [tsum_apply h_sum, tsum_apply (h_sum_a a)], }, + rw [this, lintegral_tsum (λ n, (measurable.of_uncurry_left (hf n)).ae_measurable)], + congr' with n, + rw with_density_apply' _ (hf n) a hs, +end + +/-- If a kernel `κ` is finite and a function `f : α → β → ℝ≥0∞` is bounded, then `with_density κ f` +is finite. -/ +lemma is_finite_kernel_with_density_of_bounded (κ : kernel α β) [is_finite_kernel κ] + {B : ℝ≥0∞} (hB_top : B ≠ ∞) (hf_B : ∀ a b, f a b ≤ B) : + is_finite_kernel (with_density κ f) := +begin + by_cases hf : measurable (function.uncurry f), + { exact + ⟨⟨B * is_finite_kernel.bound κ, ennreal.mul_lt_top hB_top (is_finite_kernel.bound_ne_top κ), + λ a, + begin + rw with_density_apply' κ hf a measurable_set.univ, + calc ∫⁻ b in set.univ, f a b ∂(κ a) + ≤ ∫⁻ b in set.univ, B ∂(κ a) : lintegral_mono (hf_B a) + ... = B * κ a set.univ : by simp only [measure.restrict_univ, lintegral_const] + ... ≤ B * is_finite_kernel.bound κ : + ennreal.mul_le_mul le_rfl (measure_le_bound κ a set.univ), + end⟩⟩, }, + { rw with_density_of_not_measurable _ hf, + apply_instance, }, +end + +/-- Auxiliary lemma for `is_s_finite_kernel_with_density`. +If a kernel `κ` is finite, then `with_density κ f` is s-finite. -/ +lemma is_s_finite_kernel_with_density_of_is_finite_kernel (κ : kernel α β) [is_finite_kernel κ] + (hf_ne_top : ∀ a b, f a b ≠ ∞) : + is_s_finite_kernel (with_density κ f) := +begin + -- We already have that for `f` bounded from above and a `κ` a finite kernel, + -- `with_density κ f` is finite. We write any function as a countable sum of bounded + -- functions, and decompose an s-finite kernel as a sum of finite kernels. We then use that + -- `with_density` commutes with sums for both arguments and get a sum of finite kernels. + by_cases hf : measurable (function.uncurry f), + swap, { rw with_density_of_not_measurable _ hf, apply_instance, }, + let fs : ℕ → α → β → ℝ≥0∞ := λ n a b, min (f a b) (n + 1) - min (f a b) n, + have h_le : ∀ a b n, ⌈(f a b).to_real⌉₊ ≤ n → f a b ≤ n, + { intros a b n hn, + have : (f a b).to_real ≤ n := nat.le_of_ceil_le hn, + rw ← ennreal.le_of_real_iff_to_real_le (hf_ne_top a b) _ at this, + { refine this.trans (le_of_eq _), + rw ennreal.of_real_coe_nat, }, + { norm_cast, + exact zero_le _, }, }, + have h_zero : ∀ a b n, ⌈(f a b).to_real⌉₊ ≤ n → fs n a b = 0, + { intros a b n hn, + suffices : min (f a b) (n + 1) = f a b ∧ min (f a b) n = f a b, + { simp_rw [fs, this.1, this.2, tsub_self (f a b)], }, + exact ⟨min_eq_left ((h_le a b n hn).trans (le_add_of_nonneg_right zero_le_one)), + min_eq_left (h_le a b n hn)⟩, }, + have hf_eq_tsum : f = ∑' n, fs n, + { have h_sum_a : ∀ a, summable (λ n, fs n a), + { refine λ a, pi.summable.mpr (λ b, _), + suffices : ∀ n, n ∉ finset.range ⌈(f a b).to_real⌉₊ → fs n a b = 0, + from summable_of_ne_finset_zero this, + intros n hn_not_mem, + rw [finset.mem_range, not_lt] at hn_not_mem, + exact h_zero a b n hn_not_mem, }, + ext a b : 2, + rw [tsum_apply (pi.summable.mpr h_sum_a), tsum_apply (h_sum_a a), + ennreal.tsum_eq_liminf_sum_nat], + have h_finset_sum : ∀ n, ∑ i in finset.range n, fs i a b = min (f a b) n, + { intros n, + induction n with n hn, + { simp only [finset.range_zero, finset.sum_empty, algebra_map.coe_zero, min_zero], }, + rw [finset.sum_range_succ, hn], + simp_rw [fs], + norm_cast, + rw add_tsub_cancel_iff_le, + refine min_le_min le_rfl _, + norm_cast, + exact nat.le_succ n, }, + simp_rw h_finset_sum, + refine (filter.tendsto.liminf_eq _).symm, + refine filter.tendsto.congr' _ tendsto_const_nhds, + rw [filter.eventually_eq, filter.eventually_at_top], + exact ⟨⌈(f a b).to_real⌉₊, λ n hn, (min_eq_left (h_le a b n hn)).symm⟩, }, + rw [hf_eq_tsum, with_density_tsum _ (λ (n : ℕ), _)], + swap, { exact (hf.min measurable_const).sub (hf.min measurable_const), }, + refine is_s_finite_kernel_sum (λ n, _), + suffices : is_finite_kernel (with_density κ (fs n)), by { haveI := this, apply_instance, }, + refine is_finite_kernel_with_density_of_bounded _ (ennreal.coe_ne_top : (↑n + 1) ≠ ∞) (λ a b, _), + norm_cast, + calc fs n a b ≤ min (f a b) (n + 1) : tsub_le_self + ... ≤ (n + 1) : min_le_right _ _ + ... = ↑(n + 1) : by norm_cast, +end + +/-- For a s-finite kernel `κ` and a function `f : α → β → ℝ≥0∞` which is everywhere finite, +`with_density κ f` is s-finite. -/ +theorem is_s_finite_kernel.with_density (κ : kernel α β) [is_s_finite_kernel κ] + (hf_ne_top : ∀ a b, f a b ≠ ∞) : + is_s_finite_kernel (with_density κ f) := +begin + have h_eq_sum : with_density κ f = kernel.sum (λ i, with_density (seq κ i) f), + { rw ← with_density_kernel_sum _ _, + congr, + exact (kernel_sum_seq κ).symm, }, + rw h_eq_sum, + exact is_s_finite_kernel_sum + (λ n, is_s_finite_kernel_with_density_of_is_finite_kernel (seq κ n) hf_ne_top), +end + +/-- For a s-finite kernel `κ` and a function `f : α → β → ℝ≥0`, `with_density κ f` is s-finite. -/ +instance (κ : kernel α β) [is_s_finite_kernel κ] (f : α → β → ℝ≥0) : + is_s_finite_kernel (with_density κ (λ a b, f a b)) := +is_s_finite_kernel.with_density κ (λ _ _, ennreal.coe_ne_top) + +end with_density + end kernel end probability_theory From f434b205b76aa9f9e83de0e04057a73b52352525 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 20 Feb 2023 08:13:50 +0000 Subject: [PATCH 0498/1291] chore(number_theory/sum_four_squares): squeeze simps (#18467) 17.4s -> 816ms --- src/number_theory/sum_four_squares.lean | 9 +++++---- 1 file changed, 5 insertions(+), 4 deletions(-) diff --git a/src/number_theory/sum_four_squares.lean b/src/number_theory/sum_four_squares.lean index 34800f552522d..d41635aff5967 100644 --- a/src/number_theory/sum_four_squares.lean +++ b/src/number_theory/sum_four_squares.lean @@ -90,9 +90,11 @@ let ⟨i, hσ⟩ := this (coe ∘ f) (by rw [← @zero_mul (zmod 2) _ m, ← int.cast_mul, ← h]; simp only [int.cast_add, int.cast_pow]; refl) in let σ := swap i 0 in have h01 : 2 ∣ f (σ 0) ^ 2 + f (σ 1) ^ 2, - from (char_p.int_cast_eq_zero_iff (zmod 2) 2 _).1 $ by simpa [σ] using hσ.1, + from (char_p.int_cast_eq_zero_iff (zmod 2) 2 _).1 $ + by simpa only [int.cast_pow, int.cast_add, equiv.swap_apply_right, zmod.pow_card] using hσ.1, have h23 : 2 ∣ f (σ 2) ^ 2 + f (σ 3) ^ 2, - from (char_p.int_cast_eq_zero_iff (zmod 2) 2 _).1 $ by simpa using hσ.2, + from (char_p.int_cast_eq_zero_iff (zmod 2) 2 _).1 $ + by simpa only [int.cast_pow, int.cast_add, zmod.pow_card] using hσ.2, let ⟨x, hx⟩ := h01 in let ⟨y, hy⟩ := h23 in ⟨(f (σ 0) - f (σ 1)) / 2, (f (σ 0) + f (σ 1)) / 2, (f (σ 2) - f (σ 3)) / 2, (f (σ 2) + f (σ 3)) / 2, begin @@ -101,8 +103,7 @@ let ⟨x, hx⟩ := h01 in let ⟨y, hy⟩ := h23 in ← mul_right_inj' (show (2 : ℤ) ≠ 0, from dec_trivial), ← h, mul_add, ← hx, ← hy], have : ∑ x, f (σ x)^2 = ∑ x, f x^2, { conv_rhs { rw ←equiv.sum_comp σ } }, - have fin4univ : (univ : finset (fin 4)).1 = 0 ::ₘ 1 ::ₘ 2 ::ₘ 3 ::ₘ 0, from dec_trivial, - simpa [finset.sum_eq_multiset_sum, fin4univ, multiset.sum_cons, f, add_assoc] + simpa only [fin.sum_univ_four, add_assoc] using this, end⟩ private lemma prime_sum_four_squares (p : ℕ) [hp : fact p.prime] : From 57a30493469f1a4338a5b3237b31ad8e4f3dd661 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 20 Feb 2023 09:29:42 +0000 Subject: [PATCH 0499/1291] chore(linear_algebra/prod): add a missing lemma (#18445) --- src/linear_algebra/prod.lean | 3 +++ 1 file changed, 3 insertions(+) diff --git a/src/linear_algebra/prod.lean b/src/linear_algebra/prod.lean index 98b279f10df23..83f212d8f9e1e 100644 --- a/src/linear_algebra/prod.lean +++ b/src/linear_algebra/prod.lean @@ -222,6 +222,9 @@ prod_ext_iff.2 ⟨hl, hr⟩ def prod_map (f : M →ₗ[R] M₃) (g : M₂ →ₗ[R] M₄) : (M × M₂) →ₗ[R] (M₃ × M₄) := (f.comp (fst R M M₂)).prod (g.comp (snd R M M₂)) +lemma coe_prod_map (f : M →ₗ[R] M₃) (g : M₂ →ₗ[R] M₄) : + ⇑(f.prod_map g) = prod.map f g := rfl + @[simp] theorem prod_map_apply (f : M →ₗ[R] M₃) (g : M₂ →ₗ[R] M₄) (x) : f.prod_map g x = (f x.1, g x.2) := rfl From bd9851ca476957ea4549eb19b40e7b5ade9428cc Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 20 Feb 2023 11:58:53 +0000 Subject: [PATCH 0500/1291] feat(data/prod): bijectivity of prod.map (#18446) I needed this for showing a statement about invertibility of a map built with `linear_map.prod_map` (aka a "block diagonal" linear map), where the `nonempty` conditions are trivially true. Forward-port will be at https://github.com/leanprover-community/mathlib4/pull/2346 --- src/data/prod/basic.lean | 55 +++++++++++++++++++++++++ src/data/sum/basic.lean | 26 ++++++++++++ src/number_theory/sum_four_squares.lean | 2 +- 3 files changed, 82 insertions(+), 1 deletion(-) diff --git a/src/data/prod/basic.lean b/src/data/prod/basic.lean index ca46ef41721bf..240beda2eb6e2 100644 --- a/src/data/prod/basic.lean +++ b/src/data/prod/basic.lean @@ -260,3 +260,58 @@ lemma involutive.prod_map {f : α → α} {g : β → β} : left_inverse.prod_map end function + +namespace prod +open function + +@[simp] lemma map_injective [nonempty α] [nonempty β] {f : α → γ} {g : β → δ} : + injective (map f g) ↔ injective f ∧ injective g := +⟨λ h, ⟨λ a₁ a₂ ha, begin + inhabit β, + injection @h (a₁, default) (a₂, default) (congr_arg (λ c : γ, prod.mk c (g default)) ha : _), +end, λ b₁ b₂ hb, begin + inhabit α, + injection @h (default, b₁) (default, b₂) (congr_arg (prod.mk (f default)) hb : _), +end⟩, λ h, h.1.prod_map h.2⟩ + +@[simp] lemma map_surjective [nonempty γ] [nonempty δ] {f : α → γ} {g : β → δ} : + surjective (map f g) ↔ surjective f ∧ surjective g := +⟨λ h, ⟨λ c, begin + inhabit δ, + obtain ⟨⟨a, b⟩, h⟩ := h (c, default), + exact ⟨a, congr_arg prod.fst h⟩, +end, λ d, begin + inhabit γ, + obtain ⟨⟨a, b⟩, h⟩ := h (default, d), + exact ⟨b, congr_arg prod.snd h⟩, +end⟩, λ h, h.1.prod_map h.2⟩ + +@[simp] lemma map_bijective [nonempty α] [nonempty β] {f : α → γ} {g : β → δ} : + bijective (map f g) ↔ bijective f ∧ bijective g := +begin + haveI := nonempty.map f ‹_›, + haveI := nonempty.map g ‹_›, + exact (map_injective.and map_surjective).trans (and_and_and_comm _ _ _ _) +end + +@[simp] lemma map_left_inverse [nonempty β] [nonempty δ] + {f₁ : α → β} {g₁ : γ → δ} {f₂ : β → α} {g₂ : δ → γ} : + left_inverse (map f₁ g₁) (map f₂ g₂) ↔ left_inverse f₁ f₂ ∧ left_inverse g₁ g₂ := +⟨λ h, ⟨λ b, begin + inhabit δ, + exact congr_arg prod.fst (h (b, default)), +end, λ d, begin + inhabit β, + exact congr_arg prod.snd (h (default, d)), +end⟩, λ h, h.1.prod_map h.2⟩ + +@[simp] lemma map_right_inverse [nonempty α] [nonempty γ] + {f₁ : α → β} {g₁ : γ → δ} {f₂ : β → α} {g₂ : δ → γ} : + right_inverse (map f₁ g₁) (map f₂ g₂) ↔ right_inverse f₁ f₂ ∧ right_inverse g₁ g₂ := +map_left_inverse + +@[simp] lemma map_involutive [nonempty α] [nonempty β] {f : α → α} {g : β → β} : + involutive (map f g) ↔ involutive f ∧ involutive g := +map_left_inverse + +end prod diff --git a/src/data/sum/basic.lean b/src/data/sum/basic.lean index 8436c31ff01de..bec135e6da38c 100644 --- a/src/data/sum/basic.lean +++ b/src/data/sum/basic.lean @@ -388,11 +388,37 @@ lemma surjective.sum_map {f : α → β} {g : α' → β'} (hf : surjective f) ( | (inl y) := let ⟨x, hx⟩ := hf y in ⟨inl x, congr_arg inl hx⟩ | (inr y) := let ⟨x, hx⟩ := hg y in ⟨inr x, congr_arg inr hx⟩ +lemma bijective.sum_map {f : α → β} {g : α' → β'} (hf : bijective f) (hg : bijective g) : + bijective (sum.map f g) := +⟨hf.injective.sum_map hg.injective, hf.surjective.sum_map hg.surjective⟩ + end function namespace sum open function +@[simp] lemma map_injective {f : α → γ} {g : β → δ} : + injective (sum.map f g) ↔ injective f ∧ injective g := +⟨λ h, ⟨λ a₁ a₂ ha, inl_injective $ @h (inl a₁) (inl a₂) (congr_arg inl ha : _), + λ b₁ b₂ hb, inr_injective $ @h (inr b₁) (inr b₂) (congr_arg inr hb : _)⟩, + λ h, h.1.sum_map h.2⟩ + +@[simp] lemma map_surjective {f : α → γ} {g : β → δ} : + surjective (sum.map f g) ↔ surjective f ∧ surjective g := +⟨λ h, ⟨λ c, begin + obtain ⟨a | b, h⟩ := h (inl c), + { exact ⟨a, inl_injective h⟩ }, + { cases h }, +end, λ d, begin + obtain ⟨a | b, h⟩ := h (inr d), + { cases h }, + { exact ⟨b, inr_injective h⟩ }, +end⟩, λ h, h.1.sum_map h.2⟩ + +@[simp] lemma map_bijective {f : α → γ} {g : β → δ} : + bijective (sum.map f g) ↔ bijective f ∧ bijective g := +(map_injective.and map_surjective).trans $ and_and_and_comm _ _ _ _ + lemma elim_const_const (c : γ) : sum.elim (const _ c : α → γ) (const _ c : β → γ) = const _ c := by { ext x, cases x; refl } diff --git a/src/number_theory/sum_four_squares.lean b/src/number_theory/sum_four_squares.lean index d41635aff5967..c8aee31197186 100644 --- a/src/number_theory/sum_four_squares.lean +++ b/src/number_theory/sum_four_squares.lean @@ -85,7 +85,7 @@ have ∀ f : fin 4 → zmod 2, (f 0)^2 + (f 1)^2 + (f 2)^2 + (f 3)^2 = 0 → from dec_trivial, let f : fin 4 → ℤ := vector.nth (a ::ᵥ b ::ᵥ c ::ᵥ d ::ᵥ vector.nil) in -let ⟨i, hσ⟩ := this (coe ∘ f) (by rw [← @zero_mul (zmod 2) _ m, +let ⟨i, hσ⟩ := this (λ x, coe (f x)) (by rw [← @zero_mul (zmod 2) _ m, ← show ((2 : ℤ) : zmod 2) = 0, from rfl, ← int.cast_mul, ← h]; simp only [int.cast_add, int.cast_pow]; refl) in let σ := swap i 0 in From 87c54600fe3cdc7d32ff5b50873ac724d86aef8d Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 21 Feb 2023 21:07:36 +0000 Subject: [PATCH 0501/1291] refactor(linear_algebra/bilinear_map): move lemmas about `basis` to a new file. (#18468) This cuts the import path. The copyright comes from #12269. I chose not to just move these lemmas to `linear_algebra/basis` as this file is already huge. The moved lemmas are now stated slightly more generally (`comm_semiring` rather than `comm_ring`), which seems to have been an accident in the original file as no proofs had to change. --- .../affine_space/combination.lean | 1 + src/linear_algebra/basis/bilinear.lean | 61 +++++++++++++++++++ src/linear_algebra/bilinear_map.lean | 44 ------------- src/linear_algebra/sesquilinear_form.lean | 1 + src/ring_theory/ideal/operations.lean | 1 + 5 files changed, 64 insertions(+), 44 deletions(-) create mode 100644 src/linear_algebra/basis/bilinear.lean diff --git a/src/linear_algebra/affine_space/combination.lean b/src/linear_algebra/affine_space/combination.lean index 267ad4f89860c..d2bd478803367 100644 --- a/src/linear_algebra/affine_space/combination.lean +++ b/src/linear_algebra/affine_space/combination.lean @@ -6,6 +6,7 @@ Authors: Joseph Myers import algebra.invertible import algebra.indicator_function import algebra.module.big_operators +import data.fintype.big_operators import linear_algebra.affine_space.affine_map import linear_algebra.affine_space.affine_subspace import linear_algebra.finsupp diff --git a/src/linear_algebra/basis/bilinear.lean b/src/linear_algebra/basis/bilinear.lean new file mode 100644 index 0000000000000..a5ec2cc8ffe60 --- /dev/null +++ b/src/linear_algebra/basis/bilinear.lean @@ -0,0 +1,61 @@ +/- +Copyright (c) 2022 Moritz Doll. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Moritz Doll +-/ +import linear_algebra.basis +import linear_algebra.bilinear_map + +/-! +# Lemmas about bilinear maps with a basis over each argument +-/ +namespace linear_map + +variables {ι₁ ι₂ : Type*} +variables {R R₂ S S₂ M N P : Type*} +variables {Mₗ Nₗ Pₗ : Type*} +variables [comm_semiring R] [comm_semiring S] [comm_semiring R₂] [comm_semiring S₂] + +section add_comm_monoid + +variables [add_comm_monoid M] [add_comm_monoid N] [add_comm_monoid P] +variables [add_comm_monoid Mₗ] [add_comm_monoid Nₗ] [add_comm_monoid Pₗ] +variables [module R M] [module S N] [module R₂ P] [module S₂ P] +variables [module R Mₗ] [module R Nₗ] [module R Pₗ] +variables [smul_comm_class S₂ R₂ P] +variables {ρ₁₂ : R →+* R₂} {σ₁₂ : S →+* S₂} +variables (b₁ : basis ι₁ R M) (b₂ : basis ι₂ S N) (b₁' : basis ι₁ R Mₗ) (b₂' : basis ι₂ R Nₗ) + + +/-- Two bilinear maps are equal when they are equal on all basis vectors. -/ +lemma ext_basis {B B' : M →ₛₗ[ρ₁₂] N →ₛₗ[σ₁₂] P} + (h : ∀ i j, B (b₁ i) (b₂ j) = B' (b₁ i) (b₂ j)) : B = B' := +b₁.ext $ λ i, b₂.ext $ λ j, h i j + +/-- Write out `B x y` as a sum over `B (b i) (b j)` if `b` is a basis. + +Version for semi-bilinear maps, see `sum_repr_mul_repr_mul` for the bilinear version. -/ +lemma sum_repr_mul_repr_mulₛₗ {B : M →ₛₗ[ρ₁₂] N →ₛₗ[σ₁₂] P} (x y) : + (b₁.repr x).sum (λ i xi, (b₂.repr y).sum (λ j yj, (ρ₁₂ xi) • (σ₁₂ yj) • B (b₁ i) (b₂ j))) = + B x y := +begin + conv_rhs { rw [← b₁.total_repr x, ← b₂.total_repr y] }, + simp_rw [finsupp.total_apply, finsupp.sum, map_sum₂, map_sum, + linear_map.map_smulₛₗ₂, linear_map.map_smulₛₗ], +end + +/-- Write out `B x y` as a sum over `B (b i) (b j)` if `b` is a basis. + +Version for bilinear maps, see `sum_repr_mul_repr_mulₛₗ` for the semi-bilinear version. -/ +lemma sum_repr_mul_repr_mul {B : Mₗ →ₗ[R] Nₗ →ₗ[R] Pₗ} (x y) : + (b₁'.repr x).sum (λ i xi, (b₂'.repr y).sum (λ j yj, xi • yj • B (b₁' i) (b₂' j))) = + B x y := +begin + conv_rhs { rw [← b₁'.total_repr x, ← b₂'.total_repr y] }, + simp_rw [finsupp.total_apply, finsupp.sum, map_sum₂, map_sum, + linear_map.map_smul₂, linear_map.map_smul], +end + +end add_comm_monoid + +end linear_map diff --git a/src/linear_algebra/bilinear_map.lean b/src/linear_algebra/bilinear_map.lean index 2d9681ce2e4b9..59c4b5f5fd37f 100644 --- a/src/linear_algebra/bilinear_map.lean +++ b/src/linear_algebra/bilinear_map.lean @@ -5,7 +5,6 @@ Authors: Kenny Lau, Mario Carneiro -/ import linear_algebra.basic -import linear_algebra.basis /-! # Basics on bilinear maps @@ -30,7 +29,6 @@ commuting actions, and `ρ₁₂ : R →+* R₂` and `σ₁₂ : S →+* S₂`. bilinear -/ -variables {ι₁ ι₂ : Type*} namespace linear_map @@ -311,48 +309,6 @@ variables {R R₂ S S₂ M N P : Type*} variables {Mₗ Nₗ Pₗ : Type*} variables [comm_ring R] [comm_ring S] [comm_ring R₂] [comm_ring S₂] -section add_comm_monoid - -variables [add_comm_monoid M] [add_comm_monoid N] [add_comm_monoid P] -variables [add_comm_monoid Mₗ] [add_comm_monoid Nₗ] [add_comm_monoid Pₗ] -variables [module R M] [module S N] [module R₂ P] [module S₂ P] -variables [module R Mₗ] [module R Nₗ] [module R Pₗ] -variables [smul_comm_class S₂ R₂ P] -variables {ρ₁₂ : R →+* R₂} {σ₁₂ : S →+* S₂} -variables (b₁ : basis ι₁ R M) (b₂ : basis ι₂ S N) (b₁' : basis ι₁ R Mₗ) (b₂' : basis ι₂ R Nₗ) - - -/-- Two bilinear maps are equal when they are equal on all basis vectors. -/ -lemma ext_basis {B B' : M →ₛₗ[ρ₁₂] N →ₛₗ[σ₁₂] P} - (h : ∀ i j, B (b₁ i) (b₂ j) = B' (b₁ i) (b₂ j)) : B = B' := -b₁.ext $ λ i, b₂.ext $ λ j, h i j - -/-- Write out `B x y` as a sum over `B (b i) (b j)` if `b` is a basis. - -Version for semi-bilinear maps, see `sum_repr_mul_repr_mul` for the bilinear version. -/ -lemma sum_repr_mul_repr_mulₛₗ {B : M →ₛₗ[ρ₁₂] N →ₛₗ[σ₁₂] P} (x y) : - (b₁.repr x).sum (λ i xi, (b₂.repr y).sum (λ j yj, (ρ₁₂ xi) • (σ₁₂ yj) • B (b₁ i) (b₂ j))) = - B x y := -begin - conv_rhs { rw [← b₁.total_repr x, ← b₂.total_repr y] }, - simp_rw [finsupp.total_apply, finsupp.sum, map_sum₂, map_sum, - linear_map.map_smulₛₗ₂, linear_map.map_smulₛₗ], -end - -/-- Write out `B x y` as a sum over `B (b i) (b j)` if `b` is a basis. - -Version for bilinear maps, see `sum_repr_mul_repr_mulₛₗ` for the semi-bilinear version. -/ -lemma sum_repr_mul_repr_mul {B : Mₗ →ₗ[R] Nₗ →ₗ[R] Pₗ} (x y) : - (b₁'.repr x).sum (λ i xi, (b₂'.repr y).sum (λ j yj, xi • yj • B (b₁' i) (b₂' j))) = - B x y := -begin - conv_rhs { rw [← b₁'.total_repr x, ← b₂'.total_repr y] }, - simp_rw [finsupp.total_apply, finsupp.sum, map_sum₂, map_sum, - linear_map.map_smul₂, linear_map.map_smul], -end - -end add_comm_monoid - section add_comm_group variables [add_comm_group M] [add_comm_group N] [add_comm_group P] diff --git a/src/linear_algebra/sesquilinear_form.lean b/src/linear_algebra/sesquilinear_form.lean index 088ef30ba2f0b..c7817ccaadfcb 100644 --- a/src/linear_algebra/sesquilinear_form.lean +++ b/src/linear_algebra/sesquilinear_form.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Andreas Swerdlow -/ import algebra.module.linear_map +import linear_algebra.basis.bilinear import linear_algebra.bilinear_map import algebra.euclidean_domain.instances import ring_theory.non_zero_divisors diff --git a/src/ring_theory/ideal/operations.lean b/src/ring_theory/ideal/operations.lean index 822a3ef92eeab..313ce2f53cd4c 100644 --- a/src/ring_theory/ideal/operations.lean +++ b/src/ring_theory/ideal/operations.lean @@ -6,6 +6,7 @@ Authors: Kenny Lau import algebra.algebra.operations import algebra.ring.equiv import data.nat.choose.sum +import linear_algebra.basis.bilinear import ring_theory.coprime.lemmas import ring_theory.ideal.quotient import ring_theory.non_zero_divisors From 23aa88e32dcc9d2a24cca7bc23268567ed4cd7d6 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Wed, 22 Feb 2023 02:55:41 +0000 Subject: [PATCH 0502/1291] chore(*): add mathlib4 synchronization comments (#18464) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.algebra.basic` * `algebra.algebra.hom` * `algebra.algebra.prod` * `algebra.module.algebra` * `category_theory.adjunction.basic` * `category_theory.category.ulift` * `category_theory.comm_sq` * `category_theory.conj` * `category_theory.core` * `category_theory.discrete_category` * `category_theory.monoidal.category` * `category_theory.pempty` * `category_theory.punit` * `category_theory.sums.basic` * `category_theory.yoneda` * `combinatorics.hales_jewett` * `computability.NFA` * `computability.epsilon_NFA` * `data.analysis.topology` * `data.dfinsupp.multiset` * `data.multiset.interval` * `data.pfunctor.multivariate.M` * `data.pfunctor.multivariate.W` * `data.pfunctor.multivariate.basic` * `data.pfunctor.univariate.M` * `data.qpf.multivariate.basic` * `data.qpf.multivariate.constructions.comp` * `data.qpf.multivariate.constructions.const` * `data.qpf.multivariate.constructions.prj` * `data.qpf.multivariate.constructions.quot` * `data.rat.nnrat` * `data.real.nnreal` * `deprecated.subfield` * `deprecated.subring` * `group_theory.commuting_probability` * `linear_algebra.finsupp` * `linear_algebra.prod` * `linear_algebra.quotient` * `number_theory.class_number.admissible_abs` * `order.extension.well` * `order.filter.filter_product` * `order.pfilter` * `order.prime_ideal` * `ring_theory.ideal.basic` * `set_theory.ordinal.exponential` * `set_theory.ordinal.fixed_point` * `set_theory.ordinal.natural_ops` * `set_theory.ordinal.topology` * `topology.algebra.group.basic` * `topology.algebra.group.compact` * `topology.is_locally_homeomorph` * `topology.local_homeomorph` * `topology.uniform_space.completion` --- src/algebra/algebra/basic.lean | 3 +++ src/algebra/algebra/hom.lean | 3 +++ src/algebra/algebra/prod.lean | 3 +++ src/algebra/module/algebra.lean | 3 +++ src/category_theory/adjunction/basic.lean | 3 +++ src/category_theory/category/ulift.lean | 3 +++ src/category_theory/comm_sq.lean | 3 +++ src/category_theory/conj.lean | 3 +++ src/category_theory/core.lean | 3 +++ src/category_theory/discrete_category.lean | 3 +++ src/category_theory/monoidal/category.lean | 3 +++ src/category_theory/pempty.lean | 3 +++ src/category_theory/punit.lean | 3 +++ src/category_theory/sums/basic.lean | 3 +++ src/category_theory/yoneda.lean | 3 +++ src/combinatorics/hales_jewett.lean | 3 +++ src/computability/NFA.lean | 3 +++ src/computability/epsilon_NFA.lean | 3 +++ src/data/analysis/topology.lean | 3 +++ src/data/dfinsupp/multiset.lean | 3 +++ src/data/multiset/interval.lean | 3 +++ src/data/pfunctor/multivariate/M.lean | 3 +++ src/data/pfunctor/multivariate/W.lean | 3 +++ src/data/pfunctor/multivariate/basic.lean | 3 +++ src/data/pfunctor/univariate/M.lean | 3 +++ src/data/qpf/multivariate/basic.lean | 3 +++ src/data/qpf/multivariate/constructions/comp.lean | 3 +++ src/data/qpf/multivariate/constructions/const.lean | 3 +++ src/data/qpf/multivariate/constructions/prj.lean | 3 +++ src/data/qpf/multivariate/constructions/quot.lean | 3 +++ src/data/rat/nnrat.lean | 3 +++ src/data/real/nnreal.lean | 3 +++ src/deprecated/subfield.lean | 3 +++ src/deprecated/subring.lean | 3 +++ src/group_theory/commuting_probability.lean | 3 +++ src/linear_algebra/finsupp.lean | 3 +++ src/linear_algebra/prod.lean | 3 +++ src/linear_algebra/quotient.lean | 3 +++ src/number_theory/class_number/admissible_abs.lean | 3 +++ src/order/extension/well.lean | 3 +++ src/order/filter/filter_product.lean | 3 +++ src/order/pfilter.lean | 3 +++ src/order/prime_ideal.lean | 3 +++ src/ring_theory/ideal/basic.lean | 3 +++ src/set_theory/ordinal/exponential.lean | 3 +++ src/set_theory/ordinal/fixed_point.lean | 3 +++ src/set_theory/ordinal/natural_ops.lean | 3 +++ src/set_theory/ordinal/topology.lean | 3 +++ src/topology/algebra/group/basic.lean | 3 +++ src/topology/algebra/group/compact.lean | 3 +++ src/topology/is_locally_homeomorph.lean | 3 +++ src/topology/local_homeomorph.lean | 3 +++ src/topology/uniform_space/completion.lean | 3 +++ 53 files changed, 159 insertions(+) diff --git a/src/algebra/algebra/basic.lean b/src/algebra/algebra/basic.lean index 827e2cf0bc4b5..431f754bc9471 100644 --- a/src/algebra/algebra/basic.lean +++ b/src/algebra/algebra/basic.lean @@ -17,6 +17,9 @@ import tactic.abel /-! # Algebras over commutative semirings +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define associative unital `algebra`s over commutative (semi)rings, algebra homomorphisms `alg_hom`, and algebra equivalences `alg_equiv`. diff --git a/src/algebra/algebra/hom.lean b/src/algebra/algebra/hom.lean index e0bc7c2029b1d..d2338f7e7720c 100644 --- a/src/algebra/algebra/hom.lean +++ b/src/algebra/algebra/hom.lean @@ -8,6 +8,9 @@ import algebra.algebra.basic /-! # Homomorphisms of `R`-algebras +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines bundled homomorphisms of `R`-algebras. ## Main definitions diff --git a/src/algebra/algebra/prod.lean b/src/algebra/algebra/prod.lean index 7d24d8b4c08d0..a19cc46ff8b62 100644 --- a/src/algebra/algebra/prod.lean +++ b/src/algebra/algebra/prod.lean @@ -8,6 +8,9 @@ import algebra.algebra.hom /-! # The R-algebra structure on products of R-algebras +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The R-algebra structure on `Π i : I, A i` when each `A i` is an R-algebra. ## Main defintions diff --git a/src/algebra/module/algebra.lean b/src/algebra/module/algebra.lean index 129e92cffef04..b6659e7d6af38 100644 --- a/src/algebra/module/algebra.lean +++ b/src/algebra/module/algebra.lean @@ -8,6 +8,9 @@ import algebra.algebra.basic /-! # Additional facts about modules over algebras. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ namespace linear_map diff --git a/src/category_theory/adjunction/basic.lean b/src/category_theory/adjunction/basic.lean index 0df5be9436e53..9000cc8e1b0b5 100644 --- a/src/category_theory/adjunction/basic.lean +++ b/src/category_theory/adjunction/basic.lean @@ -8,6 +8,9 @@ import category_theory.equivalence /-! # Adjunctions between functors +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + `F ⊣ G` represents the data of an adjunction between two functors `F : C ⥤ D` and `G : D ⥤ C`. `F` is the left adjoint and `G` is the right adjoint. diff --git a/src/category_theory/category/ulift.lean b/src/category_theory/category/ulift.lean index f9ee205cd0503..f1ee02a5c9c96 100644 --- a/src/category_theory/category/ulift.lean +++ b/src/category_theory/category/ulift.lean @@ -10,6 +10,9 @@ import category_theory.eq_to_hom /-! # Basic API for ulift +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains a very basic API for working with the categorical instance on `ulift C` where `C` is a type with a category instance. diff --git a/src/category_theory/comm_sq.lean b/src/category_theory/comm_sq.lean index 19e70cbc27cbe..7e9e2ad8bd048 100644 --- a/src/category_theory/comm_sq.lean +++ b/src/category_theory/comm_sq.lean @@ -8,6 +8,9 @@ import category_theory.arrow /-! # Commutative squares +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file provide an API for commutative squares in categories. If `top`, `left`, `right` and `bottom` are four morphisms which are the edges of a square, `comm_sq top left right bottom` is the predicate that this diff --git a/src/category_theory/conj.lean b/src/category_theory/conj.lean index d2baceff27987..0b16aede4368d 100644 --- a/src/category_theory/conj.lean +++ b/src/category_theory/conj.lean @@ -9,6 +9,9 @@ import category_theory.endomorphism /-! # Conjugate morphisms by isomorphisms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + An isomorphism `α : X ≅ Y` defines - a monoid isomorphism `conj : End X ≃* End Y` by `α.conj f = α.inv ≫ f ≫ α.hom`; - a group isomorphism `conj_Aut : Aut X ≃* Aut Y` by `α.conj_Aut f = α.symm ≪≫ f ≪≫ α`. diff --git a/src/category_theory/core.lean b/src/category_theory/core.lean index 8c55b96351640..2311c8fe3d085 100644 --- a/src/category_theory/core.lean +++ b/src/category_theory/core.lean @@ -11,6 +11,9 @@ import category_theory.types /-! # The core of a category +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The core of a category `C` is the (non-full) subcategory of `C` consisting of all objects, and all isomorphisms. We construct it as a `groupoid`. diff --git a/src/category_theory/discrete_category.lean b/src/category_theory/discrete_category.lean index f3c0a11643369..fb4f2e90d6442 100644 --- a/src/category_theory/discrete_category.lean +++ b/src/category_theory/discrete_category.lean @@ -9,6 +9,9 @@ import data.ulift /-! # Discrete categories +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define `discrete α` as a structure containing a term `a : α` for any type `α`, and use this type alias to provide a `small_category` instance whose only morphisms are the identities. diff --git a/src/category_theory/monoidal/category.lean b/src/category_theory/monoidal/category.lean index 08f64db141167..d475a9e130525 100644 --- a/src/category_theory/monoidal/category.lean +++ b/src/category_theory/monoidal/category.lean @@ -8,6 +8,9 @@ import category_theory.products.basic /-! # Monoidal categories +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A monoidal category is a category equipped with a tensor product, unitors, and an associator. In the definition, we provide the tensor product as a pair of functions * `tensor_obj : C → C → C` diff --git a/src/category_theory/pempty.lean b/src/category_theory/pempty.lean index 9188797a5e7b2..003d021913812 100644 --- a/src/category_theory/pempty.lean +++ b/src/category_theory/pempty.lean @@ -8,6 +8,9 @@ import category_theory.discrete_category /-! # The empty category +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Defines a category structure on `pempty`, and the unique functor `pempty ⥤ C` for any category `C`. -/ diff --git a/src/category_theory/punit.lean b/src/category_theory/punit.lean index 86c16356489a6..a953c39ac8dc1 100644 --- a/src/category_theory/punit.lean +++ b/src/category_theory/punit.lean @@ -9,6 +9,9 @@ import category_theory.discrete_category /-! # The category `discrete punit` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define `star : C ⥤ discrete punit` sending everything to `punit.star`, show that any two functors to `discrete punit` are naturally isomorphic, and construct the equivalence `(discrete punit ⥤ C) ≌ C`. diff --git a/src/category_theory/sums/basic.lean b/src/category_theory/sums/basic.lean index d994cf0f200d6..06ba51c49750d 100644 --- a/src/category_theory/sums/basic.lean +++ b/src/category_theory/sums/basic.lean @@ -8,6 +8,9 @@ import category_theory.eq_to_hom /-! # Binary disjoint unions of categories +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define the category instance on `C ⊕ D` when `C` and `D` are categories. We define: diff --git a/src/category_theory/yoneda.lean b/src/category_theory/yoneda.lean index daebdd43a83bb..4d6761f2974e8 100644 --- a/src/category_theory/yoneda.lean +++ b/src/category_theory/yoneda.lean @@ -10,6 +10,9 @@ import category_theory.products.basic /-! # The Yoneda embedding +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The Yoneda embedding as a functor `yoneda : C ⥤ (Cᵒᵖ ⥤ Type v₁)`, along with an instance that it is `fully_faithful`. diff --git a/src/combinatorics/hales_jewett.lean b/src/combinatorics/hales_jewett.lean index 05db18003681c..f7677fc09aca0 100644 --- a/src/combinatorics/hales_jewett.lean +++ b/src/combinatorics/hales_jewett.lean @@ -11,6 +11,9 @@ import algebra.big_operators.basic /-! # The Hales-Jewett theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We prove the Hales-Jewett theorem and deduce Van der Waerden's theorem as a corollary. The Hales-Jewett theorem is a result in Ramsey theory dealing with *combinatorial lines*. Given diff --git a/src/computability/NFA.lean b/src/computability/NFA.lean index b1fce95851172..1026ac008c6fb 100644 --- a/src/computability/NFA.lean +++ b/src/computability/NFA.lean @@ -8,6 +8,9 @@ import data.fintype.powerset /-! # Nondeterministic Finite Automata + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. This file contains the definition of a Nondeterministic Finite Automaton (NFA), a state machine which determines whether a string (implemented as a list over an arbitrary alphabet) is in a regular set by evaluating the string over every possible path. diff --git a/src/computability/epsilon_NFA.lean b/src/computability/epsilon_NFA.lean index 6faa55eafe884..6035d3fa6977d 100644 --- a/src/computability/epsilon_NFA.lean +++ b/src/computability/epsilon_NFA.lean @@ -9,6 +9,9 @@ import computability.NFA /-! # Epsilon Nondeterministic Finite Automata +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains the definition of an epsilon Nondeterministic Finite Automaton (`ε_NFA`), a state machine which determines whether a string (implemented as a list over an arbitrary alphabet) is in a regular set by evaluating the string over every possible path, also having access to ε-transitons, diff --git a/src/data/analysis/topology.lean b/src/data/analysis/topology.lean index e3e37b4705d22..c76fd7b2ac460 100644 --- a/src/data/analysis/topology.lean +++ b/src/data/analysis/topology.lean @@ -9,6 +9,9 @@ import topology.bases /-! # Computational realization of topological spaces (experimental) +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file provides infrastructure to compute with topological spaces. ## Main declarations diff --git a/src/data/dfinsupp/multiset.lean b/src/data/dfinsupp/multiset.lean index ade908fd4639b..de23e647e818c 100644 --- a/src/data/dfinsupp/multiset.lean +++ b/src/data/dfinsupp/multiset.lean @@ -8,6 +8,9 @@ import data.dfinsupp.order /-! # Equivalence between `multiset` and `ℕ`-valued finitely supported functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This defines `dfinsupp.to_multiset` the equivalence between `Π₀ a : α, ℕ` and `multiset α`, along with `multiset.to_dfinsupp` the reverse equivalence. diff --git a/src/data/multiset/interval.lean b/src/data/multiset/interval.lean index ce6951ea6a794..33ae1fb363183 100644 --- a/src/data/multiset/interval.lean +++ b/src/data/multiset/interval.lean @@ -11,6 +11,9 @@ import data.nat.interval /-! # Finite intervals of multisets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file provides the `locally_finite_order` instance for `multiset α` and calculates the cardinality of its finite intervals. diff --git a/src/data/pfunctor/multivariate/M.lean b/src/data/pfunctor/multivariate/M.lean index ede48f428f644..b74d0ef8d645a 100644 --- a/src/data/pfunctor/multivariate/M.lean +++ b/src/data/pfunctor/multivariate/M.lean @@ -9,6 +9,9 @@ import data.pfunctor.univariate.M /-! # The M construction as a multivariate polynomial functor. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + M types are potentially infinite tree-like structures. They are defined as the greatest fixpoint of a polynomial functor. diff --git a/src/data/pfunctor/multivariate/W.lean b/src/data/pfunctor/multivariate/W.lean index 3923df2b0c67f..59dc86e7a3cd6 100644 --- a/src/data/pfunctor/multivariate/W.lean +++ b/src/data/pfunctor/multivariate/W.lean @@ -8,6 +8,9 @@ import data.pfunctor.multivariate.basic /-! # The W construction as a multivariate polynomial functor. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + W types are well-founded tree-like structures. They are defined as the least fixpoint of a polynomial functor. diff --git a/src/data/pfunctor/multivariate/basic.lean b/src/data/pfunctor/multivariate/basic.lean index 1e0dacfdec7d7..3a4a950f74b7b 100644 --- a/src/data/pfunctor/multivariate/basic.lean +++ b/src/data/pfunctor/multivariate/basic.lean @@ -9,6 +9,9 @@ import data.pfunctor.univariate.basic /-! # Multivariate polynomial functors. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Multivariate polynomial functors are used for defining M-types and W-types. They map a type vector `α` to the type `Σ a : A, B a ⟹ α`, with `A : Type` and `B : A → typevec n`. They interact well with Lean's inductive definitions because diff --git a/src/data/pfunctor/univariate/M.lean b/src/data/pfunctor/univariate/M.lean index 8a2be6295b3f8..3090f1d746cf8 100644 --- a/src/data/pfunctor/univariate/M.lean +++ b/src/data/pfunctor/univariate/M.lean @@ -8,6 +8,9 @@ import data.pfunctor.univariate.basic /-! # M-types +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + M types are potentially infinite tree-like structures. They are defined as the greatest fixpoint of a polynomial functor. -/ diff --git a/src/data/qpf/multivariate/basic.lean b/src/data/qpf/multivariate/basic.lean index 6b47237bf5943..3aba3258e6ac9 100644 --- a/src/data/qpf/multivariate/basic.lean +++ b/src/data/qpf/multivariate/basic.lean @@ -8,6 +8,9 @@ import data.pfunctor.multivariate.basic /-! # Multivariate quotients of polynomial functors. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Basic definition of multivariate QPF. QPFs form a compositional framework for defining inductive and coinductive types, their quotients and nesting. diff --git a/src/data/qpf/multivariate/constructions/comp.lean b/src/data/qpf/multivariate/constructions/comp.lean index 00e56d65ee4c5..2d5b9260590d3 100644 --- a/src/data/qpf/multivariate/constructions/comp.lean +++ b/src/data/qpf/multivariate/constructions/comp.lean @@ -10,6 +10,9 @@ import data.qpf.multivariate.basic /-! # The composition of QPFs is itself a QPF +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define composition between one `n`-ary functor and `n` `m`-ary functors and show that it preserves the QPF structure -/ diff --git a/src/data/qpf/multivariate/constructions/const.lean b/src/data/qpf/multivariate/constructions/const.lean index 113adf4a31bb3..f3c9b395698d4 100644 --- a/src/data/qpf/multivariate/constructions/const.lean +++ b/src/data/qpf/multivariate/constructions/const.lean @@ -10,6 +10,9 @@ import data.qpf.multivariate.basic /-! # Constant functors are QPFs +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Constant functors map every type vectors to the same target type. This is a useful device for constructing data types from more basic types that are not actually functorial. For instance `const n nat` makes diff --git a/src/data/qpf/multivariate/constructions/prj.lean b/src/data/qpf/multivariate/constructions/prj.lean index e27a1be1f8700..959e3d1617b7c 100644 --- a/src/data/qpf/multivariate/constructions/prj.lean +++ b/src/data/qpf/multivariate/constructions/prj.lean @@ -8,6 +8,9 @@ import control.functor.multivariate import data.qpf.multivariate.basic /-! +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Projection functors are QPFs. The `n`-ary projection functors on `i` is an `n`-ary functor `F` such that `F (α₀..αᵢ₋₁, αᵢ, αᵢ₊₁..αₙ₋₁) = αᵢ` -/ diff --git a/src/data/qpf/multivariate/constructions/quot.lean b/src/data/qpf/multivariate/constructions/quot.lean index d85cbbf011d42..0869d40073f28 100644 --- a/src/data/qpf/multivariate/constructions/quot.lean +++ b/src/data/qpf/multivariate/constructions/quot.lean @@ -9,6 +9,9 @@ import data.qpf.multivariate.basic /-! # The quotient of QPF is itself a QPF +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The quotients are here defined using a surjective function and its right inverse. They are very similar to the `abs` and `repr` functions found in the definition of `mvqpf` diff --git a/src/data/rat/nnrat.lean b/src/data/rat/nnrat.lean index e6c1963ed1fb4..87cc872ebb68e 100644 --- a/src/data/rat/nnrat.lean +++ b/src/data/rat/nnrat.lean @@ -9,6 +9,9 @@ import algebra.order.nonneg.field /-! # Nonnegative rationals +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the nonnegative rationals as a subtype of `rat` and provides its algebraic order structure. diff --git a/src/data/real/nnreal.lean b/src/data/real/nnreal.lean index 15685b7067bba..43084ff20997a 100644 --- a/src/data/real/nnreal.lean +++ b/src/data/real/nnreal.lean @@ -13,6 +13,9 @@ import tactic.positivity /-! # Nonnegative real numbers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define `nnreal` (notation: `ℝ≥0`) to be the type of non-negative real numbers, a.k.a. the interval `[0, ∞)`. We also define the following operations and structures on `ℝ≥0`: diff --git a/src/deprecated/subfield.lean b/src/deprecated/subfield.lean index 45d8eec408488..2c68a30c862e5 100644 --- a/src/deprecated/subfield.lean +++ b/src/deprecated/subfield.lean @@ -8,6 +8,9 @@ import deprecated.subring /-! # Unbundled subfields (deprecated) +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file is deprecated, and is no longer imported by anything in mathlib other than other deprecated files, and test files. You should not need to import it. diff --git a/src/deprecated/subring.lean b/src/deprecated/subring.lean index 3c88103c3f8b8..fa0f809eb5fae 100644 --- a/src/deprecated/subring.lean +++ b/src/deprecated/subring.lean @@ -10,6 +10,9 @@ import ring_theory.subring.basic /-! # Unbundled subrings (deprecated) +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file is deprecated, and is no longer imported by anything in mathlib other than other deprecated files, and test files. You should not need to import it. diff --git a/src/group_theory/commuting_probability.lean b/src/group_theory/commuting_probability.lean index 67a21e7762b7b..bc035f0065209 100644 --- a/src/group_theory/commuting_probability.lean +++ b/src/group_theory/commuting_probability.lean @@ -11,6 +11,9 @@ import group_theory.index /-! # Commuting Probability + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. This file introduces the commuting probability of finite groups. ## Main definitions diff --git a/src/linear_algebra/finsupp.lean b/src/linear_algebra/finsupp.lean index f4acf825b69fe..8e38ee6586f15 100644 --- a/src/linear_algebra/finsupp.lean +++ b/src/linear_algebra/finsupp.lean @@ -10,6 +10,9 @@ import linear_algebra.span /-! # Properties of the module `α →₀ M` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given an `R`-module `M`, the `R`-module structure on `α →₀ M` is defined in `data.finsupp.basic`. diff --git a/src/linear_algebra/prod.lean b/src/linear_algebra/prod.lean index 83f212d8f9e1e..2b190f6abffca 100644 --- a/src/linear_algebra/prod.lean +++ b/src/linear_algebra/prod.lean @@ -9,6 +9,9 @@ import algebra.algebra.prod /-! ### Products of modules +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines constructors for linear maps whose domains or codomains are products. It contains theorems relating these to each other, as well as to `submodule.prod`, `submodule.map`, diff --git a/src/linear_algebra/quotient.lean b/src/linear_algebra/quotient.lean index 2e91ca51d780d..e76a86d7e1d91 100644 --- a/src/linear_algebra/quotient.lean +++ b/src/linear_algebra/quotient.lean @@ -9,6 +9,9 @@ import linear_algebra.span /-! # Quotients by submodules +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + * If `p` is a submodule of `M`, `M ⧸ p` is the quotient of `M` with respect to `p`: that is, elements of `M` are identified if their difference is in `p`. This is itself a module. diff --git a/src/number_theory/class_number/admissible_abs.lean b/src/number_theory/class_number/admissible_abs.lean index 53bff2817009c..b93949539f369 100644 --- a/src/number_theory/class_number/admissible_abs.lean +++ b/src/number_theory/class_number/admissible_abs.lean @@ -8,6 +8,9 @@ import number_theory.class_number.admissible_absolute_value /-! # Admissible absolute value on the integers + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. This file defines an admissible absolute value `absolute_value.abs_is_admissible` which we use to show the class number of the ring of integers of a number field is finite. diff --git a/src/order/extension/well.lean b/src/order/extension/well.lean index b684cf4c87f44..2aeeb6c1b5e94 100644 --- a/src/order/extension/well.lean +++ b/src/order/extension/well.lean @@ -9,6 +9,9 @@ import set_theory.ordinal.arithmetic /-! # Extend a well-founded order to a well-order +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file constructs a well-order (linear well-founded order) which is an extension of a given well-founded order. diff --git a/src/order/filter/filter_product.lean b/src/order/filter/filter_product.lean index e2c823a2f90d4..1f7d0867a1d18 100644 --- a/src/order/filter/filter_product.lean +++ b/src/order/filter/filter_product.lean @@ -9,6 +9,9 @@ import order.filter.germ /-! # Ultraproducts +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + If `φ` is an ultrafilter, then the space of germs of functions `f : α → β` at `φ` is called the *ultraproduct*. In this file we prove properties of ultraproducts that rely on `φ` being an ultrafilter. Definitions and properties that work for any filter should go to `order.filter.germ`. diff --git a/src/order/pfilter.lean b/src/order/pfilter.lean index 33fa6d787f53a..e9ddf29995b33 100644 --- a/src/order/pfilter.lean +++ b/src/order/pfilter.lean @@ -8,6 +8,9 @@ import order.ideal /-! # Order filters +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions Throughout this file, `P` is at least a preorder, but some sections diff --git a/src/order/prime_ideal.lean b/src/order/prime_ideal.lean index d113dabce9f8e..1ca320f66c408 100644 --- a/src/order/prime_ideal.lean +++ b/src/order/prime_ideal.lean @@ -9,6 +9,9 @@ import order.pfilter /-! # Prime ideals +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions Throughout this file, `P` is at least a preorder, but some sections require more diff --git a/src/ring_theory/ideal/basic.lean b/src/ring_theory/ideal/basic.lean index c84f67ef4343f..21f966e5aafed 100644 --- a/src/ring_theory/ideal/basic.lean +++ b/src/ring_theory/ideal/basic.lean @@ -14,6 +14,9 @@ import linear_algebra.finsupp # Ideals over a ring +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + 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. diff --git a/src/set_theory/ordinal/exponential.lean b/src/set_theory/ordinal/exponential.lean index ac04319279e5e..4b210acb5a2fe 100644 --- a/src/set_theory/ordinal/exponential.lean +++ b/src/set_theory/ordinal/exponential.lean @@ -7,6 +7,9 @@ import set_theory.ordinal.arithmetic /-! # Ordinal exponential +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define the power function and the logarithm function on ordinals, -/ diff --git a/src/set_theory/ordinal/fixed_point.lean b/src/set_theory/ordinal/fixed_point.lean index e59d33952814a..f0b504b61d1e8 100644 --- a/src/set_theory/ordinal/fixed_point.lean +++ b/src/set_theory/ordinal/fixed_point.lean @@ -10,6 +10,9 @@ import set_theory.ordinal.exponential /-! # Fixed points of normal functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We prove various statements about the fixed points of normal ordinal functions. We state them in three forms: as statements about type-indexed families of normal functions, as statements about ordinal-indexed families of normal functions, and as statements about a single normal function. For diff --git a/src/set_theory/ordinal/natural_ops.lean b/src/set_theory/ordinal/natural_ops.lean index 526bc5ea3d49c..ff6f1e849beff 100644 --- a/src/set_theory/ordinal/natural_ops.lean +++ b/src/set_theory/ordinal/natural_ops.lean @@ -9,6 +9,9 @@ import set_theory.ordinal.arithmetic /-! # Natural operations on ordinals +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The goal of this file is to define natural addition and multiplication on ordinals, also known as the Hessenberg sum and product, and provide a basic API. The natural addition of two ordinals `a ♯ b` is recursively defined as the least ordinal greater than `a' ♯ b` and `a ♯ b'` for `a' < a` diff --git a/src/set_theory/ordinal/topology.lean b/src/set_theory/ordinal/topology.lean index 3936f9e539058..329858582a348 100644 --- a/src/set_theory/ordinal/topology.lean +++ b/src/set_theory/ordinal/topology.lean @@ -9,6 +9,9 @@ import topology.order.basic /-! ### Topology of ordinals +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We prove some miscellaneous results involving the order topology of ordinals. ### Main results diff --git a/src/topology/algebra/group/basic.lean b/src/topology/algebra/group/basic.lean index 5b5dd64a855bd..a70983e5cfdd1 100644 --- a/src/topology/algebra/group/basic.lean +++ b/src/topology/algebra/group/basic.lean @@ -12,6 +12,9 @@ import topology.algebra.constructions /-! # Topological groups +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the following typeclasses: * `topological_group`, `topological_add_group`: multiplicative and additive topological groups, diff --git a/src/topology/algebra/group/compact.lean b/src/topology/algebra/group/compact.lean index 62ce628e6263f..2e99caa1361a7 100644 --- a/src/topology/algebra/group/compact.lean +++ b/src/topology/algebra/group/compact.lean @@ -10,6 +10,9 @@ import topology.sets.compacts /-! # Additional results on topological groups +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Two results on topological groups that have been separated out as they require more substantial imports developing either positive compacts or the compact open topology. diff --git a/src/topology/is_locally_homeomorph.lean b/src/topology/is_locally_homeomorph.lean index cdd16b8b411b7..75fa4fe4ab8ec 100644 --- a/src/topology/is_locally_homeomorph.lean +++ b/src/topology/is_locally_homeomorph.lean @@ -8,6 +8,9 @@ import topology.local_homeomorph /-! # Local homeomorphisms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines local homeomorphisms. ## Main definitions diff --git a/src/topology/local_homeomorph.lean b/src/topology/local_homeomorph.lean index 21de0e24902e4..64ab36280df56 100644 --- a/src/topology/local_homeomorph.lean +++ b/src/topology/local_homeomorph.lean @@ -9,6 +9,9 @@ import topology.sets.opens /-! # Local homeomorphisms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines homeomorphisms between open subsets of topological spaces. An element `e` of `local_homeomorph α β` is an extension of `local_equiv α β`, i.e., it is a pair of functions `e.to_fun` and `e.inv_fun`, inverse of each other on the sets `e.source` and `e.target`. diff --git a/src/topology/uniform_space/completion.lean b/src/topology/uniform_space/completion.lean index f954d08426b0b..32df076ed4101 100644 --- a/src/topology/uniform_space/completion.lean +++ b/src/topology/uniform_space/completion.lean @@ -8,6 +8,9 @@ import topology.uniform_space.abstract_completion /-! # Hausdorff completions of uniform spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The goal is to construct a left-adjoint to the inclusion of complete Hausdorff uniform spaces into all uniform spaces. Any uniform space `α` gets a completion `completion α` and a morphism (ie. uniformly continuous map) `coe : α → completion α` which solves the universal From 464e1ddc4026874cdea576f9d698df42779d31c9 Mon Sep 17 00:00:00 2001 From: negiizhao Date: Wed, 22 Feb 2023 11:26:43 +0000 Subject: [PATCH 0503/1291] feat(data/real/ereal): positive covariance typeclasses (#18157) There is no semiring structure on `ereal`, but we can prove results about monotonicity of multiplication by positive elements. This PR adds some lemmas and typeclass `pos_mul_mono` `mul_pos_mono` `pos_mul_reflect_lt` `mul_pos_reflect_lt` for `ereal`. Co-authored-by: Eric Wieser --- src/data/real/ereal.lean | 112 +++++++++++++++++++++++++++++++-------- 1 file changed, 89 insertions(+), 23 deletions(-) diff --git a/src/data/real/ereal.lean b/src/data/real/ereal.lean index 67cf654e91ce0..a3a5bf5e7f728 100644 --- a/src/data/real/ereal.lean +++ b/src/data/real/ereal.lean @@ -54,7 +54,7 @@ noncomputable theory /-- ereal : The type `[-∞, ∞]` -/ @[derive [has_bot, has_zero, has_one, nontrivial, add_monoid, - has_Sup, has_Inf, complete_linear_order, linear_ordered_add_comm_monoid]] + has_Sup, has_Inf, complete_linear_order, linear_ordered_add_comm_monoid, zero_le_one_class]] def ereal := with_bot (with_top ℝ) /-- The canonical inclusion froms reals to ereals. Do not use directly: as this is registered as @@ -281,6 +281,10 @@ end | ⊤ := rfl | (some x) := rfl +@[simp] lemma coe_ennreal_of_real {x : ℝ} : + (ennreal.of_real x : ereal) = max x 0 := +rfl + lemma coe_nnreal_eq_coe_real (x : ℝ≥0) : ((x : ℝ≥0∞) : ereal) = (x : ℝ) := rfl @[simp, norm_cast] lemma coe_ennreal_zero : ((0 : ℝ≥0∞) : ereal) = 0 := rfl @@ -715,6 +719,37 @@ begin { simp only [top_mul_top, to_real_top, mul_zero] } end +protected lemma neg_mul (x y : ereal) : -x * y = -(x * y) := +begin + induction x using ereal.rec; induction y using ereal.rec, + { refl }, + { rcases lt_trichotomy 0 y with hy | rfl | hy, + { rw [bot_mul_coe_of_pos hy, neg_bot, top_mul_coe_of_pos hy] }, + { rw [coe_zero, mul_zero, mul_zero, neg_zero] }, + { rw [bot_mul_coe_of_neg hy, neg_bot, neg_top, top_mul_coe_of_neg hy] } }, + { refl }, + { rcases lt_trichotomy 0 x with hx | rfl | hx, + { rw [coe_mul_bot_of_pos hx, neg_bot, ← coe_neg, coe_mul_bot_of_neg (neg_neg_of_pos hx)], }, + { rw [coe_zero, zero_mul, neg_zero, zero_mul] }, + { rw [coe_mul_bot_of_neg hx, neg_top, ← coe_neg, coe_mul_bot_of_pos (neg_pos_of_neg hx)], }, }, + { norm_cast, exact neg_mul _ _, }, + { rcases lt_trichotomy 0 x with hx | rfl | hx, + { rw [coe_mul_top_of_pos hx, neg_top, ← coe_neg, coe_mul_top_of_neg (neg_neg_of_pos hx)], }, + { rw [coe_zero, zero_mul, neg_zero, zero_mul] }, + { rw [coe_mul_top_of_neg hx, neg_bot, ← coe_neg, coe_mul_top_of_pos (neg_pos_of_neg hx)], }, }, + { refl }, + { rcases lt_trichotomy 0 y with hy | rfl | hy, + { rw [top_mul_coe_of_pos hy, neg_top, bot_mul_coe_of_pos hy] }, + { rw [coe_zero, mul_zero, mul_zero, neg_zero] }, + { rw [top_mul_coe_of_neg hy, neg_top, neg_bot, bot_mul_coe_of_neg hy] } }, + { refl } +end + +instance : has_distrib_neg ereal := +{ neg_mul := ereal.neg_mul, + mul_neg := λ x y, by { rw [x.mul_comm, x.mul_comm], exact y.neg_mul x, }, + ..ereal.has_involutive_neg } + /-! ### Absolute value -/ /-- The absolute value from `ereal` to `ℝ≥0∞`, mapping `⊥` and `⊤` to `⊤` and @@ -727,6 +762,8 @@ protected def abs : ereal → ℝ≥0∞ @[simp] lemma abs_top : (⊤ : ereal).abs = ⊤ := rfl @[simp] lemma abs_bot : (⊥ : ereal).abs = ⊤ := rfl +lemma abs_def (x : ℝ) : (x : ereal).abs = ennreal.of_real (|x|) := rfl + lemma abs_coe_lt_top (x : ℝ) : (x : ereal).abs < ⊤ := ennreal.of_real_lt_top @@ -741,6 +778,9 @@ end @[simp] lemma abs_zero : (0 : ereal).abs = 0 := by rw [abs_eq_zero_iff] +@[simp] lemma coe_abs (x : ℝ) : ((x : ereal).abs : ereal) = (|x| : ℝ) := +by rcases lt_trichotomy 0 x with hx | rfl | hx; simp [abs_def] + @[simp] lemma abs_mul (x y : ereal) : (x * y).abs = x.abs * y.abs := begin symmetry, @@ -814,31 +854,43 @@ begin { refl } end +lemma sign_mul_abs (x : ereal) : + (sign x * x.abs : ereal) = x := +begin + induction x using ereal.rec, + { simp }, + { rcases lt_trichotomy 0 x with hx | rfl | hx, + { simp [sign_pos hx, abs_of_pos hx] }, + { simp }, + { simp [sign_neg hx, abs_of_neg hx] } }, + { simp } +end + lemma sign_eq_and_abs_eq_iff_eq {x y : ereal} : (x.abs = y.abs ∧ sign x = sign y) ↔ x = y := begin - split, swap, - { rintros rfl, simp only [eq_self_iff_true, and_self] }, - rintros ⟨habs, hsign⟩, - induction x using ereal.rec; induction y using ereal.rec, - { refl }, - { simpa only using abs_coe_lt_top y }, - { simpa only using hsign }, - { simpa only using abs_coe_lt_top x }, - { have : |x| = |y|, - by simpa only [ereal.abs, ennreal.of_real_eq_of_real_iff, abs_nonneg] using habs, - rcases abs_eq_abs.1 this with rfl|h, - { refl }, - { rcases lt_trichotomy x 0 with hx|rfl|hx, - { have hy : 0 < y, by simpa only [h, right.neg_neg_iff] using hx, - simpa only [hx, hy, sign_coe, sign_neg, sign_pos] using hsign }, - { simp only [zero_eq_neg.1 h] }, - { have hy : y < 0, by simpa only [h, right.neg_pos_iff] using hx, - simpa only [hx, hy, sign_coe, sign_neg, sign_pos] using hsign } } }, - { simpa only using abs_coe_lt_top x }, - { simpa only }, - { simpa only using abs_coe_lt_top y }, - { refl } + split, + { rintros ⟨habs, hsign⟩, rw [← x.sign_mul_abs, ← y.sign_mul_abs, habs, hsign] }, + { rintros rfl, simp only [eq_self_iff_true, and_self] } +end + +lemma le_iff_sign {x y : ereal} : + x ≤ y ↔ sign x < sign y ∨ + sign x = sign_type.neg ∧ sign y = sign_type.neg ∧ y.abs ≤ x.abs ∨ + sign x = sign_type.zero ∧ sign y = sign_type.zero ∨ + sign x = sign_type.pos ∧ sign y = sign_type.pos ∧ x.abs ≤ y.abs := +begin + split, + { intro h, + rcases (sign.monotone h).lt_or_eq with hs | hs, + { exact or.inl hs }, + { rw [← x.sign_mul_abs, ← y.sign_mul_abs] at h, + cases sign y; rw [hs] at *, + { simp }, + { simp at ⊢ h, exact or.inl h }, + { simpa using h, }, }, }, + { rintros (h | h | h | h), { exact (sign.monotone.reflect_lt h).le, }, + all_goals { rw [← x.sign_mul_abs, ← y.sign_mul_abs], simp [h] } } end instance : comm_monoid_with_zero ereal := @@ -849,6 +901,20 @@ instance : comm_monoid_with_zero ereal := mul_comm := ereal.mul_comm, ..ereal.has_mul, ..ereal.has_one, ..ereal.has_zero, ..ereal.mul_zero_one_class } +instance : pos_mul_mono ereal := +pos_mul_mono_iff_covariant_pos.2 ⟨begin + rintros ⟨x, x0⟩ a b h, dsimp, + rcases le_iff_sign.mp h with h | h | h | h, + { rw [le_iff_sign], left, simp [sign_pos x0, h] }, + all_goals { rw [← x.sign_mul_abs, ← a.sign_mul_abs, ← b.sign_mul_abs, sign_pos x0], + simp only [h], dsimp, + simp only [neg_mul, mul_neg, ereal.neg_le_neg_iff, one_mul, le_refl, zero_mul, mul_zero] }, + all_goals { norm_cast, exact mul_le_mul_left' h.2.2 _, }, +end⟩ +instance : mul_pos_mono ereal := pos_mul_mono_iff_mul_pos_mono.1 ereal.pos_mul_mono +instance : pos_mul_reflect_lt ereal := pos_mul_mono.to_pos_mul_reflect_lt +instance : mul_pos_reflect_lt ereal := mul_pos_mono.to_mul_pos_reflect_lt + @[simp, norm_cast] lemma coe_pow (x : ℝ) (n : ℕ) : (↑(x ^ n) : ereal) = x ^ n := map_pow (⟨coe, coe_one, coe_mul⟩ : ℝ →* ereal) _ _ From b875cbb7f2aa2b4c685aaa2f99705689c95322ad Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 22 Feb 2023 12:50:51 +0000 Subject: [PATCH 0504/1291] =?UTF-8?q?feat(linear=5Falgebra/affine=5Fspace/?= =?UTF-8?q?affine=5Fsubspace):=20`=E2=8A=A5=20<=20affine=5Fspan=20k=20s=20?= =?UTF-8?q?=E2=86=94=20s.nonempty`=20(#18170)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Two simple corollaries. Also make `s : set P` implicit in the existing lemmas as it can be inferred from the rewrite. --- .../affine_space/affine_subspace.lean | 23 ++++++++++++------- .../affine_space/finite_dimensional.lean | 6 ++--- 2 files changed, 18 insertions(+), 11 deletions(-) diff --git a/src/linear_algebra/affine_space/affine_subspace.lean b/src/linear_algebra/affine_space/affine_subspace.lean index d7e5ebcb29fe2..bb0b5eddee5f1 100644 --- a/src/linear_algebra/affine_space/affine_subspace.lean +++ b/src/linear_algebra/affine_space/affine_subspace.lean @@ -1121,22 +1121,29 @@ begin { exact λ ⟨i₁, hi₁, hv⟩, ⟨p i₁, ⟨i₁, ⟨set.mem_univ _, hi₁⟩, rfl⟩, hv⟩ } end -/-- The affine span of a set is nonempty if and only if that set -is. -/ -lemma affine_span_nonempty (s : set P) : - (affine_span k s : set P).nonempty ↔ s.nonempty := +section +variables {s : set P} + +/-- The affine span of a set is nonempty if and only if that set is. -/ +lemma affine_span_nonempty : (affine_span k s : set P).nonempty ↔ s.nonempty := span_points_nonempty k s +alias affine_span_nonempty ↔ _ _root_.set.nonempty.affine_span + /-- The affine span of a nonempty set is nonempty. -/ -instance {s : set P} [nonempty s] : nonempty (affine_span k s) := -((affine_span_nonempty k s).mpr (nonempty_subtype.mp ‹_›)).to_subtype +instance [nonempty s] : nonempty (affine_span k s) := +((nonempty_coe_sort.1 ‹_›).affine_span _).to_subtype /-- The affine span of a set is `⊥` if and only if that set is empty. -/ -@[simp] lemma affine_span_eq_bot {s : set P} : - affine_span k s = ⊥ ↔ s = ∅ := +@[simp] lemma affine_span_eq_bot : affine_span k s = ⊥ ↔ s = ∅ := by rw [←not_iff_not, ←ne.def, ←ne.def, ←nonempty_iff_ne_bot, affine_span_nonempty, nonempty_iff_ne_empty] +@[simp] lemma bot_lt_affine_span : ⊥ < affine_span k s ↔ s.nonempty := +by { rw [bot_lt_iff_ne_bot, nonempty_iff_ne_empty], exact (affine_span_eq_bot _).not } + +end + variables {k} /-- diff --git a/src/linear_algebra/affine_space/finite_dimensional.lean b/src/linear_algebra/affine_space/finite_dimensional.lean index a4934d582cc60..27419bb618522 100644 --- a/src/linear_algebra/affine_space/finite_dimensional.lean +++ b/src/linear_algebra/affine_space/finite_dimensional.lean @@ -227,9 +227,9 @@ lemma affine_independent.affine_span_image_finset_eq_of_le_of_card_eq_finrank_ad [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 hn : s.nonempty, + { rw [←finset.card_pos, hc], apply nat.succ_pos }, + refine eq_of_direction_eq_of_nonempty_of_le _ ((hn.image _).to_set.affine_span _)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 From 4550138052d0a416b700c27056d492e2ef53214e Mon Sep 17 00:00:00 2001 From: Jon Eugster Date: Wed, 22 Feb 2023 15:42:26 +0000 Subject: [PATCH 0505/1291] feat(data/set): add lemmas about `powerset` of `insert` and `singleton` (#18356) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Add lemmas about `𝒫 {x}` and `𝒫 (insert x A)`. Mathlib4 pair : leanprover-community/mathlib4/pull/2000 Co-authored-by: Eric Wieser --- src/data/set/basic.lean | 20 ++++++++++++++++++++ src/data/set/image.lean | 23 +++++++++++++++++++++++ 2 files changed, 43 insertions(+) diff --git a/src/data/set/basic.lean b/src/data/set/basic.lean index 661e500e3b9cf..729ea6cb1e1db 100644 --- a/src/data/set/basic.lean +++ b/src/data/set/basic.lean @@ -698,6 +698,22 @@ begin exacts [(ha hx).elim, hxt] end +theorem subset_insert_iff_of_not_mem {s t : set α} {a : α} (h : a ∉ s) : + s ⊆ insert a t ↔ s ⊆ t := +begin + constructor, + { intros g y hy, + specialize g hy, + rw [mem_insert_iff] at g, + rcases g with g | g, + { rw [g] at hy, + contradiction }, + { assumption }}, + { intros g y hy, + specialize g hy, + exact mem_insert_of_mem _ g } +end + theorem ssubset_iff_insert {s t : set α} : s ⊂ t ↔ ∃ a ∉ s, insert a s ⊆ t := begin simp only [insert_subset, exists_and_distrib_right, ssubset_def, not_subset], @@ -1301,6 +1317,10 @@ ext $ λ s, subset_empty_iff @[simp] theorem powerset_univ : 𝒫 (univ : set α) = univ := eq_univ_of_forall subset_univ +/- The powerset of a singleton contains only `∅` and the singleton itself. -/ +theorem powerset_singleton (x : α) : 𝒫 ({x} : set α) = {∅, {x}} := +by { ext y, rw [mem_powerset_iff, subset_singleton_iff_eq, mem_insert_iff, mem_singleton_iff] } + /-! ### Sets defined as an if-then-else -/ lemma mem_dite_univ_right (p : Prop) [decidable p] (t : p → set α) (x : α) : diff --git a/src/data/set/image.lean b/src/data/set/image.lean index b5d416816794a..fba4f00b8d959 100644 --- a/src/data/set/image.lean +++ b/src/data/set/image.lean @@ -471,6 +471,29 @@ end end image +/-! ### Lemmas about the powerset and image. -/ + +/-- The powerset of `{a} ∪ s` is `𝒫 s` together with `{a} ∪ t` for each `t ∈ 𝒫 s`. -/ +theorem powerset_insert (s : set α) (a : α) : + 𝒫 (insert a s) = 𝒫 s ∪ (insert a '' 𝒫 s) := +begin + ext t, + simp_rw [mem_union, mem_image, mem_powerset_iff], + split, + { intro h, + by_cases hs : a ∈ t, + { right, + refine ⟨t \ {a}, _, _⟩, + { rw [diff_singleton_subset_iff], + assumption }, + { rw [insert_diff_singleton, insert_eq_of_mem hs] }}, + { left, + exact (subset_insert_iff_of_not_mem hs).mp h}}, + { rintros (h | ⟨s', h₁, rfl⟩), + { exact subset_trans h (subset_insert a s) }, + { exact insert_subset_insert h₁ }} +end + /-! ### Lemmas about range of a function. -/ section range variables {f : ι → α} {s t : set α} From 335232c774b3d0513ab1531582779dc25d6fdc9a Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 22 Feb 2023 18:57:01 +0000 Subject: [PATCH 0506/1291] feat(analysis/normed): `mul_opposite X` inherits the same norm as `X`. (#18470) This adds a `normed_*` instance for `mul_opposite` everywhere that we also have one for `prod` and `pi`. The intention is to only do this for the additive case. It's not clear to me whether it's sensible to put the same norm on _multiplicative_ normed groups. --- src/analysis/normed/field/basic.lean | 21 +++++++++++++++++ src/analysis/normed/group/basic.lean | 34 ++++++++++++++++++++++++++++ src/analysis/normed_space/basic.lean | 9 ++++++++ 3 files changed, 64 insertions(+) diff --git a/src/analysis/normed/field/basic.lean b/src/analysis/normed/field/basic.lean index 5a559308339d4..884facaf95090 100644 --- a/src/analysis/normed/field/basic.lean +++ b/src/analysis/normed/field/basic.lean @@ -140,6 +140,10 @@ instance pi.norm_one_class {ι : Type*} {α : ι → Type*} [nonempty ι] [finty norm_one_class (Π i, α i) := ⟨by simp [pi.norm_def, finset.sup_const finset.univ_nonempty]⟩ +instance mul_opposite.norm_one_class [seminormed_add_comm_group α] [has_one α] [norm_one_class α] : + norm_one_class αᵐᵒᵖ := +⟨@norm_one α _ _ _⟩ + section non_unital_semi_normed_ring variables [non_unital_semi_normed_ring α] @@ -211,6 +215,11 @@ instance pi.non_unital_semi_normed_ring {π : ι → Type*} [fintype ι] finset.sup_mul_le_mul_sup_of_nonneg _ (λ i _, zero_le _) (λ i _, zero_le _), ..pi.seminormed_add_comm_group } +instance mul_opposite.non_unital_semi_normed_ring : non_unital_semi_normed_ring αᵐᵒᵖ := +{ norm_mul := mul_opposite.rec $ λ x, mul_opposite.rec $ λ y, + (norm_mul_le y x).trans_eq (mul_comm _ _), + ..mul_opposite.seminormed_add_comm_group } + end non_unital_semi_normed_ring section semi_normed_ring @@ -326,6 +335,10 @@ instance pi.semi_normed_ring {π : ι → Type*} [fintype ι] { ..pi.non_unital_semi_normed_ring, ..pi.seminormed_add_comm_group, } +instance mul_opposite.semi_normed_ring : semi_normed_ring αᵐᵒᵖ := +{ ..mul_opposite.non_unital_semi_normed_ring, + ..mul_opposite.seminormed_add_comm_group } + end semi_normed_ring section non_unital_normed_ring @@ -349,6 +362,10 @@ instance pi.non_unital_normed_ring {π : ι → Type*} [fintype ι] [Π i, non_u { norm_mul := norm_mul_le, ..pi.normed_add_comm_group } +instance mul_opposite.non_unital_normed_ring : non_unital_normed_ring αᵐᵒᵖ := +{ norm_mul := norm_mul_le, + ..mul_opposite.normed_add_comm_group } + end non_unital_normed_ring section normed_ring @@ -376,6 +393,10 @@ instance pi.normed_ring {π : ι → Type*} [fintype ι] [Π i, normed_ring (π { norm_mul := norm_mul_le, ..pi.normed_add_comm_group } +instance mul_opposite.normed_ring : normed_ring αᵐᵒᵖ := +{ norm_mul := norm_mul_le, + ..mul_opposite.normed_add_comm_group } + end normed_ring @[priority 100] -- see Note [lower instance priority] diff --git a/src/analysis/normed/group/basic.lean b/src/analysis/normed/group/basic.lean index 075c92a279765..273451169b6d2 100644 --- a/src/analysis/normed/group/basic.lean +++ b/src/analysis/normed/group/basic.lean @@ -1689,6 +1689,40 @@ instance pi.normed_comm_group [Π i, normed_comm_group (π i)] : normed_comm_gro end pi +/-! ### Multiplicative opposite -/ + +namespace mul_opposite + +/-- The (additive) norm on the multiplicative opposite is the same as the norm on the original type. + +Note that we do not provide this more generally as `has_norm Eᵐᵒᵖ`, as this is not always a good +choice of norm in the multiplicative `seminormed_group E` case. + +We could repeat this instance to provide a `[seminormed_group E] : seminormed_group Eᵃᵒᵖ` instance, +but that case would likely never be used. +-/ +instance [seminormed_add_group E] : seminormed_add_group Eᵐᵒᵖ := +{ norm := λ x, ‖x.unop‖, + dist_eq := λ _ _, dist_eq_norm _ _, + to_pseudo_metric_space := mul_opposite.pseudo_metric_space } + +lemma norm_op [seminormed_add_group E] (a : E) : ‖mul_opposite.op a‖ = ‖a‖ := rfl +lemma norm_unop [seminormed_add_group E] (a : Eᵐᵒᵖ) : ‖mul_opposite.unop a‖ = ‖a‖ := rfl + +lemma nnnorm_op [seminormed_add_group E] (a : E) : ‖mul_opposite.op a‖₊ = ‖a‖₊ := rfl +lemma nnnorm_unop [seminormed_add_group E] (a : Eᵐᵒᵖ) : ‖mul_opposite.unop a‖₊ = ‖a‖₊ := rfl + +instance [normed_add_group E] : normed_add_group Eᵐᵒᵖ := +{ .. mul_opposite.seminormed_add_group } + +instance [seminormed_add_comm_group E] : seminormed_add_comm_group Eᵐᵒᵖ := +{ dist_eq := λ _ _, dist_eq_norm _ _ } + +instance [normed_add_comm_group E] : normed_add_comm_group Eᵐᵒᵖ := +{ .. mul_opposite.seminormed_add_comm_group } + +end mul_opposite + /-! ### Subgroups of normed groups -/ namespace subgroup diff --git a/src/analysis/normed_space/basic.lean b/src/analysis/normed_space/basic.lean index adadf553ae992..9e6b5b018cfd7 100644 --- a/src/analysis/normed_space/basic.lean +++ b/src/analysis/normed_space/basic.lean @@ -248,6 +248,11 @@ instance pi.normed_space {E : ι → Type*} [fintype ι] [∀i, seminormed_add_c ‖a‖₊ * ↑(finset.sup finset.univ (λ (b : ι), ‖f b‖₊)), by simp only [(nnreal.coe_mul _ _).symm, nnreal.mul_finset_sup, nnnorm_smul] } +instance mul_opposite.normed_space : normed_space α Eᵐᵒᵖ := +{ norm_smul_le := λ s x, (norm_smul s x.unop).le, + ..mul_opposite.normed_add_comm_group, + ..mul_opposite.module _ } + /-- A subspace of a normed space is also a normed space, with the restriction of the norm. -/ instance submodule.normed_space {𝕜 R : Type*} [has_smul 𝕜 R] [normed_field 𝕜] [ring R] {E : Type*} [seminormed_add_comm_group E] [normed_space 𝕜 E] [module R E] @@ -519,6 +524,10 @@ instance pi.normed_algebra {E : ι → Type*} [fintype ι] { .. pi.normed_space, .. pi.algebra _ E } +instance mul_opposite.normed_algebra {E : Type*} [semi_normed_ring E] [normed_algebra 𝕜 E] : + normed_algebra 𝕜 Eᵐᵒᵖ := +{ ..mul_opposite.normed_space } + end normed_algebra /-- A non-unital algebra homomorphism from an `algebra` to a `normed_algebra` induces a From 1bda4fc53de6ade5ab9da36f2192e24e2084a2ce Mon Sep 17 00:00:00 2001 From: Chris Hughes <33847686+ChrisHughes24@users.noreply.github.com> Date: Wed, 22 Feb 2023 18:57:02 +0000 Subject: [PATCH 0507/1291] fix(group_theory/nielsen_schreier): fix typo (#18483) --- src/group_theory/nielsen_schreier.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/group_theory/nielsen_schreier.lean b/src/group_theory/nielsen_schreier.lean index aa42a81781acd..a54d8f0d7a83b 100644 --- a/src/group_theory/nielsen_schreier.lean +++ b/src/group_theory/nielsen_schreier.lean @@ -88,7 +88,7 @@ lemma ext_functor {G} [groupoid.{v} G] [is_free_groupoid G] {X : Type v} [group let ⟨_, _, u⟩ := @unique_lift G _ _ X _ (λ (a b : generators G) (e : a ⟶ b), g.map (of e)) in trans (u _ h) (u _ (λ _ _ _, rfl)).symm -/-- An action groupoid over a free froup is free. More generally, one could show that the groupoid +/-- An action groupoid over a free group is free. More generally, one could show that the groupoid of elements over a free groupoid is free, but this version is easier to prove and suffices for our purposes. From 3ade05ac9447ae31a22d2ea5423435e054131240 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 22 Feb 2023 20:57:27 +0000 Subject: [PATCH 0508/1291] feat(docs): allow deeper nesting within undergrad.yaml (#18407) Actually rendering this extra level needs https://github.com/leanprover-community/leanprover-community.github.io/pull/311 --- docs/undergrad.yaml | 9 +++++++-- scripts/yaml_check.py | 28 ++++++++++++++++++++-------- 2 files changed, 27 insertions(+), 10 deletions(-) diff --git a/docs/undergrad.yaml b/docs/undergrad.yaml index 450ce4b9a2365..134f166f3b276 100644 --- a/docs/undergrad.yaml +++ b/docs/undergrad.yaml @@ -369,8 +369,13 @@ Single Variable Complex Analysis: differentiability with respect to the complex variable: 'has_fpower_series_on_ball.differentiable_on' antiderivative: '' complex exponential: 'complex.exp' - extension of trigonometric functions to the complex plane: 'complex.sin' - power series expansion of elementary functions: 'complex.has_sum_cos' + extension of trigonometric functions to the complex plane: + cos: 'complex.cos' + sin: 'complex.sin' + power series expansion of elementary functions: + cos: 'complex.has_sum_cos' + sin: 'complex.has_sum_sin' + log: '' Functions on one complex variable: holomorphic functions: 'differentiable_on' Cauchy-Riemann conditions: '' diff --git a/scripts/yaml_check.py b/scripts/yaml_check.py index 04a07ea84fea4..5f1d1634231b7 100644 --- a/scripts/yaml_check.py +++ b/scripts/yaml_check.py @@ -1,18 +1,25 @@ -from typing import Dict, Optional, Tuple, List +from typing import Dict, Optional, Union, Tuple, List import yaml import sys -def tiered_extract(db: Dict[str, Dict[str, Dict[str, Optional[str]]]]) -> List[Tuple[str, str]]: - """From a three-level deep nested dictionary, return a list of (key, values) +TieredDict = Dict[str, Union[Optional[str], 'TieredDict']] + +def tiered_extract(db: TieredDict) -> List[Tuple[List[str], str]]: + """From a nested dictionary, return a list of (key_path, values) of the deepest level.""" out = [] - for entry in db.values(): - for values in entry.values(): - for name, decl in values.items(): - if decl and '/' not in decl: - out.append((name, decl)) + for name, entry in db.items(): + if isinstance(entry, dict): + for subname, value in tiered_extract(entry): + out.append(([name] + subname, value)) + else: + if entry and '/' not in entry: + out.append(([name], entry)) return out +def flatten_names(data: List[Tuple[List[str], str]]) -> List[Tuple[str, str]]: + return [(' :: '.join(id), v) for id, v in data] + def print_list(fn: str, pairs: List[Tuple[str, str]]) -> None: with open(fn, 'w') as out: for (id, val) in pairs: @@ -38,7 +45,12 @@ def print_list(fn: str, pairs: List[Tuple[str, str]]) -> None: hundred_decls = hundred_decls + [(index, d) for d in entry['decls']] overview_decls = tiered_extract(overview) +assert all(len(n) == 3 for n, _ in overview_decls) +overview_decls = flatten_names(overview_decls) + undergrad_decls = tiered_extract(undergrad) +assert all(len(n) >= 3 for n, _ in undergrad_decls) +undergrad_decls = flatten_names(undergrad_decls) print_list('100.txt', hundred_decls) print_list('overview.txt', overview_decls) From 6efec6bb9fcaed3cf1baaddb2eaadd8a2a06679c Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Wed, 22 Feb 2023 22:12:59 +0000 Subject: [PATCH 0509/1291] feat(topology/continuous_function): some lemmas spun off from #18392 (#18465) A few miscellaneous lemmas about continuous functions needed for my Poisson summation project. Companion mathlib4 PR (blank for now) at [#2369](https://github.com/leanprover-community/mathlib4/pull/2369). --- src/topology/algebra/monoid.lean | 21 +++++ src/topology/continuous_function/algebra.lean | 91 ++++++++++++++++--- src/topology/continuous_function/basic.lean | 8 +- src/topology/continuous_function/bounded.lean | 13 ++- src/topology/continuous_function/compact.lean | 15 ++- 5 files changed, 130 insertions(+), 18 deletions(-) diff --git a/src/topology/algebra/monoid.lean b/src/topology/algebra/monoid.lean index fa433e34f9574..5dbee1af0e363 100644 --- a/src/topology/algebra/monoid.lean +++ b/src/topology/algebra/monoid.lean @@ -7,6 +7,7 @@ import algebra.big_operators.finprod import order.filter.pointwise import topology.algebra.mul_action import algebra.big_operators.pi +import topology.continuous_function.basic /-! # Theory of topological monoids @@ -688,3 +689,23 @@ by { rw ← Inf_range, exact has_continuous_mul_Inf (set.forall_range_iff.mpr h' by { rw inf_eq_infi, refine has_continuous_mul_infi (λ b, _), cases b; assumption } end lattice_ops + +namespace continuous_map + +variables [has_mul X] [has_continuous_mul X] + +/-- The continuous map `λ y, y * x` -/ +@[to_additive "The continuous map `λ y, y + x"] +protected def mul_right (x : X) : C(X, X) := mk _ (continuous_mul_right x) + +@[to_additive, simp] +lemma coe_mul_right (x : X) : ⇑(continuous_map.mul_right x) = λ y, y * x := rfl + +/-- The continuous map `λ y, x * y` -/ +@[to_additive "The continuous map `λ y, x + y"] +protected def mul_left (x : X) : C(X, X) := mk _ (continuous_mul_left x) + +@[to_additive, simp] +lemma coe_mul_left (x : X) : ⇑(continuous_map.mul_left x) = λ y, x * y := rfl + +end continuous_map diff --git a/src/topology/continuous_function/algebra.lean b/src/topology/continuous_function/algebra.lean index 0a1bf7d10a6b7..29d3335e29e17 100644 --- a/src/topology/continuous_function/algebra.lean +++ b/src/topology/continuous_function/algebra.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison, Nicolò Cavalleri -/ import algebra.algebra.pi +import algebra.periodic import algebra.algebra.subalgebra.basic import algebra.star.star_alg_hom import tactic.field_simp @@ -45,6 +46,8 @@ namespace continuous_map variables {α : Type*} {β : Type*} {γ : Type*} variables [topological_space α] [topological_space β] [topological_space γ] +/- ### "mul" and "add" -/ + @[to_additive] instance has_mul [has_mul β] [has_continuous_mul β] : has_mul C(α, β) := ⟨λ f g, ⟨f * g, continuous_mul.comp (f.continuous.prod_mk g.continuous : _)⟩⟩ @@ -52,24 +55,34 @@ instance has_mul [has_mul β] [has_continuous_mul β] : has_mul C(α, β) := @[simp, norm_cast, to_additive] lemma coe_mul [has_mul β] [has_continuous_mul β] (f g : C(α, β)) : ⇑(f * g) = f * g := rfl +@[simp, to_additive] +lemma mul_apply [has_mul β] [has_continuous_mul β] (f g : C(α, β)) (x : α) : + (f * g) x = f x * g x := rfl + @[simp, to_additive] lemma mul_comp [has_mul γ] [has_continuous_mul γ] (f₁ f₂ : C(β, γ)) (g : C(α, β)) : (f₁ * f₂).comp g = f₁.comp g * f₂.comp g := rfl -@[to_additive] -instance [has_one β] : has_one C(α, β) := ⟨const α 1⟩ +/- ### "one" -/ -@[simp, norm_cast, to_additive] -lemma coe_one [has_one β] : ⇑(1 : C(α, β)) = 1 := rfl +@[to_additive] instance [has_one β] : has_one C(α, β) := ⟨const α 1⟩ + +@[simp, norm_cast, to_additive] lemma coe_one [has_one β] : ⇑(1 : C(α, β)) = 1 := rfl + +@[simp, to_additive] lemma one_apply [has_one β] (x : α) : (1 : C(α, β)) x = 1 := rfl @[simp, to_additive] lemma one_comp [has_one γ] (g : C(α, β)) : (1 : C(β, γ)).comp g = 1 := rfl -instance [has_nat_cast β] : has_nat_cast C(α, β) := -⟨λ n, continuous_map.const _ n⟩ +/- ### "nat_cast" -/ -@[simp, norm_cast] -lemma coe_nat_cast [has_nat_cast β] (n : ℕ) : ((n : C(α, β)) : α → β) = n := rfl +instance [has_nat_cast β] : has_nat_cast C(α, β) := ⟨λ n, continuous_map.const _ n⟩ + +@[simp, norm_cast] lemma coe_nat_cast [has_nat_cast β] (n : ℕ) : ((n : C(α, β)) : α → β) = n := rfl + +@[simp] lemma nat_cast_apply [has_nat_cast β] (n : ℕ) (x : α) : (n : C(α, β)) x = n := rfl + +/- ### "int_cast" -/ instance [has_int_cast β] : has_int_cast C(α, β) := ⟨λ n, continuous_map.const _ n⟩ @@ -77,6 +90,10 @@ instance [has_int_cast β] : has_int_cast C(α, β) := @[simp, norm_cast] lemma coe_int_cast [has_int_cast β] (n : ℤ) : ((n : C(α, β)) : α → β) = n := rfl +@[simp] lemma int_cast_apply [has_int_cast β] (n : ℤ) (x : α) : (n : C(α, β)) x = n := rfl + +/- ### "nsmul" and "pow" -/ + instance has_nsmul [add_monoid β] [has_continuous_add β] : has_smul ℕ C(α, β) := ⟨λ n f, ⟨n • f, f.continuous.nsmul n⟩⟩ @@ -88,8 +105,14 @@ instance has_pow [monoid β] [has_continuous_mul β] : has_pow C(α, β) ℕ := lemma coe_pow [monoid β] [has_continuous_mul β] (f : C(α, β)) (n : ℕ) : ⇑(f ^ n) = f ^ n := rfl --- don't make `coe_nsmul` simp as the linter complains it's redundant WRT `coe_smul` -attribute [simp] coe_pow +@[to_additive] lemma pow_apply [monoid β] [has_continuous_mul β] + (f : C(α, β)) (n : ℕ) (x : α) : + (f ^ n) x = f x ^ n := +rfl + +-- don't make auto-generated `coe_nsmul` and `nsmul_apply` simp, as the linter complains they're +-- redundant WRT `coe_smul` +attribute [simp] coe_pow pow_apply @[to_additive] lemma pow_comp [monoid γ] [has_continuous_mul γ] (f : C(β, γ)) (n : ℕ) (g : C(α, β)) : @@ -99,6 +122,8 @@ rfl -- don't make `nsmul_comp` simp as the linter complains it's redundant WRT `smul_comp` attribute [simp] pow_comp +/- ### "inv" and "neg" -/ + @[to_additive] instance [group β] [topological_group β] : has_inv C(α, β) := { inv := λ f, ⟨f⁻¹, f.continuous.inv⟩ } @@ -108,10 +133,16 @@ lemma coe_inv [group β] [topological_group β] (f : C(α, β)) : ⇑(f⁻¹) = f⁻¹ := rfl +@[simp, to_additive] lemma inv_apply [group β] [topological_group β] (f : C(α, β)) (x : α) : + f⁻¹ x = (f x)⁻¹ := +rfl + @[simp, to_additive] lemma inv_comp [group γ] [topological_group γ] (f : C(β, γ)) (g : C(α, β)) : (f⁻¹).comp g = (f.comp g)⁻¹ := rfl +/- ### "div" and "sub" -/ + @[to_additive] instance [has_div β] [has_continuous_div β] : has_div C(α, β) := { div := λ f g, ⟨f / g, f.continuous.div' g.continuous⟩ } @@ -120,11 +151,17 @@ instance [has_div β] [has_continuous_div β] : has_div C(α, β) := lemma coe_div [has_div β] [has_continuous_div β] (f g : C(α, β)) : ⇑(f / g) = f / g := rfl +@[simp, to_additive] lemma div_apply [has_div β] [has_continuous_div β] (f g : C(α, β)) (x : α) : + (f / g) x = f x / g x := +rfl + @[simp, to_additive] lemma div_comp [has_div γ] [has_continuous_div γ] (f g : C(β, γ)) (h : C(α, β)) : (f / g).comp h = (f.comp h) / (g.comp h) := rfl +/- ### "zpow" and "zsmul" -/ + instance has_zsmul [add_group β] [topological_add_group β] : has_smul ℤ C(α, β) := { smul := λ z f, ⟨z • f, f.continuous.zsmul z⟩ } @@ -138,8 +175,14 @@ lemma coe_zpow [group β] [topological_group β] (f : C(α, β)) (z : ℤ) : ⇑(f ^ z) = f ^ z := rfl --- don't make `coe_zsmul` simp as the linter complains it's redundant WRT `coe_smul` -attribute [simp] coe_zpow +@[to_additive] lemma zpow_apply [group β] [topological_group β] + (f : C(α, β)) (z : ℤ) (x : α) : + (f ^ z) x = f x ^ z := +rfl + +-- don't make auto-generated `coe_zsmul` and `zsmul_apply` simp as the linter complains they're +-- redundant WRT `coe_smul` +attribute [simp] coe_zpow zpow_apply @[to_additive] lemma zpow_comp [group γ] [topological_group γ] (f : C(β, γ)) (z : ℤ) (g : C(α, β)) : @@ -881,6 +924,30 @@ lemma comp_star_alg_hom'_comp (g : C(Y, Z)) (f : C(X, Y)) : comp_star_alg_hom' 𝕜 A (g.comp f) = (comp_star_alg_hom' 𝕜 A f).comp (comp_star_alg_hom' 𝕜 A g) := star_alg_hom.ext $ λ _, continuous_map.ext $ λ _, rfl +section periodicity + +/-! ### Summing translates of a function -/ + +/-- Summing the translates of `f` by `ℤ • p` gives a map which is periodic with period `p`. +(This is true without any convergence conditions, since if the sum doesn't converge it is taken to +be the zero map, which is periodic.) -/ +lemma periodic_tsum_comp_add_zsmul [locally_compact_space X] [add_comm_group X] + [topological_add_group X] [add_comm_monoid Y] [has_continuous_add Y] [t2_space Y] + (f : C(X, Y)) (p : X) : + function.periodic ⇑(∑' (n : ℤ), f.comp (continuous_map.add_right (n • p))) p := +begin + intro x, + by_cases h : summable (λ n : ℤ, f.comp (continuous_map.add_right (n • p))), + { convert congr_arg (λ f : C(X, Y), f x) ((equiv.add_right (1 : ℤ)).tsum_eq _) using 1, + simp_rw [←tsum_apply h, ←tsum_apply ((equiv.add_right (1 : ℤ)).summable_iff.mpr h), + equiv.coe_add_right, comp_apply, coe_add_right, add_one_zsmul, add_comm (_ • p) p, + ←add_assoc] }, + { rw tsum_eq_zero_of_not_summable h, + simp only [coe_zero, pi.zero_apply] } +end + +end periodicity + end continuous_map namespace homeomorph diff --git a/src/topology/continuous_function/basic.lean b/src/topology/continuous_function/basic.lean index 89c54869af349..897344cee1864 100644 --- a/src/topology/continuous_function/basic.lean +++ b/src/topology/continuous_function/basic.lean @@ -224,7 +224,13 @@ def restrict (f : C(α, β)) : C(s, β) := ⟨f ∘ coe⟩ @[simp] lemma coe_restrict (f : C(α, β)) : ⇑(f.restrict s) = f ∘ coe := rfl -/-- The restriction of a continuous map onto the preimage of a set. -/ +@[simp] lemma restrict_apply (f : C(α, β)) (s : set α) (x : s) : f.restrict s x = f x := rfl + +@[simp] lemma restrict_apply_mk (f : C(α, β)) (s : set α) (x : α) (hx : x ∈ s) : + f.restrict s ⟨x, hx⟩ = f x := +rfl + +/-- The restriction of a continuous map to the preimage of a set. -/ @[simps] def restrict_preimage (f : C(α, β)) (s : set β) : C(f ⁻¹' s, s) := ⟨s.restrict_preimage f, continuous_iff_continuous_at.mpr $ λ x, f.2.continuous_at.restrict_preimage⟩ diff --git a/src/topology/continuous_function/bounded.lean b/src/topology/continuous_function/bounded.lean index 10b2f516edc71..da3985efeeb04 100644 --- a/src/topology/continuous_function/bounded.lean +++ b/src/topology/continuous_function/bounded.lean @@ -308,11 +308,17 @@ begin end /-- Composition of a bounded continuous function and a continuous function. -/ -@[simps { fully_applied := ff }] def comp_continuous {δ : Type*} [topological_space δ] (f : α →ᵇ β) (g : C(δ, α)) : δ →ᵇ β := { to_continuous_map := f.1.comp g, map_bounded' := f.map_bounded'.imp (λ C hC x y, hC _ _) } +@[simp] lemma coe_comp_continuous {δ : Type*} [topological_space δ] (f : α →ᵇ β) (g : C(δ, α)) : + coe_fn (f.comp_continuous g) = f ∘ g := rfl + +@[simp] lemma comp_continuous_apply {δ : Type*} [topological_space δ] + (f : α →ᵇ β) (g : C(δ, α)) (x : δ) : f.comp_continuous g x = f (g x) := +rfl + lemma lipschitz_comp_continuous {δ : Type*} [topological_space δ] (g : C(δ, α)) : lipschitz_with 1 (λ f : α →ᵇ β, f.comp_continuous g) := lipschitz_with.mk_one $ λ f₁ f₂, (dist_le dist_nonneg).2 $ λ x, dist_coe_le_dist (g x) @@ -322,10 +328,13 @@ lemma continuous_comp_continuous {δ : Type*} [topological_space δ] (g : C(δ, (lipschitz_comp_continuous g).continuous /-- Restrict a bounded continuous function to a set. -/ -@[simps apply { fully_applied := ff }] def restrict (f : α →ᵇ β) (s : set α) : s →ᵇ β := f.comp_continuous $ (continuous_map.id _).restrict s +@[simp] lemma coe_restrict (f : α →ᵇ β) (s : set α) : coe_fn (f.restrict s) = f ∘ coe := rfl + +@[simp] lemma restrict_apply (f : α →ᵇ β) (s : set α) (x : s) : f.restrict s x = f x := rfl + /-- Composition (in the target) of a bounded continuous function with a Lipschitz map again gives a bounded continuous function -/ def comp (G : β → γ) {C : ℝ≥0} (H : lipschitz_with C G) diff --git a/src/topology/continuous_function/compact.lean b/src/topology/continuous_function/compact.lean index c20644015b441..d6c6eff9e0520 100644 --- a/src/topology/continuous_function/compact.lean +++ b/src/topology/continuous_function/compact.lean @@ -206,6 +206,11 @@ le_trans (neg_le_neg (f.norm_coe_le_norm x)) (neg_le.mp (neg_le_abs_self (f x))) lemma norm_eq_supr_norm : ‖f‖ = ⨆ x : α, ‖f x‖ := (mk_of_compact f).norm_eq_supr_norm +lemma norm_restrict_mono_set {X : Type*} [topological_space X] + (f : C(X, E)) {K L : topological_space.compacts X} (hKL : K ≤ L) : + ‖f.restrict K‖ ≤ ‖f.restrict L‖ := +(norm_le _ (norm_nonneg _)).mpr (λ x, norm_coe_le_norm (f.restrict L) $ set.inclusion hKL x) + end section @@ -409,7 +414,12 @@ map_continuous (comp_right_continuous_map A f) end comp_right -section weierstrass +section local_normal_convergence +/-! ### Local normal convergence + +A sum of continuous functions (on a locally compact space) is "locally normally convergent" if the +sum of its sup-norms on any compact subset is summable. This implies convergence in the topology +of `C(X, E)` (i.e. locally uniform convergence). -/ open topological_space @@ -427,8 +437,7 @@ begin simpa only [has_sum, A] using summable_of_summable_norm (hF K) end -end weierstrass - +end local_normal_convergence /-! ### Star structures From 3dadefa3f544b1db6214777fe47910739b54c66a Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Thu, 23 Feb 2023 00:36:51 +0000 Subject: [PATCH 0510/1291] chore(*): add mathlib4 synchronization comments (#18485) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `category_theory.adjunction.fully_faithful` * `category_theory.adjunction.whiskering` * `category_theory.category.Cat` * `category_theory.category.galois_connection` * `category_theory.fin_category` * `category_theory.isomorphism_classes` * `category_theory.lifting_properties.adjunction` * `category_theory.lifting_properties.basic` * `category_theory.limits.shapes.strong_epi` * `category_theory.skeletal` * `combinatorics.additive.pluennecke_ruzsa` * `data.qpf.multivariate.constructions.sigma` * `field_theory.subfield` * `group_theory.order_of_element` * `linear_algebra.projection` * `set_theory.ordinal.principal` * `topology.algebra.order.group` * `topology.continuous_function.ordered` --- src/category_theory/adjunction/fully_faithful.lean | 3 +++ src/category_theory/adjunction/whiskering.lean | 3 +++ src/category_theory/category/Cat.lean | 3 +++ src/category_theory/category/galois_connection.lean | 3 +++ src/category_theory/fin_category.lean | 3 +++ src/category_theory/isomorphism_classes.lean | 3 +++ src/category_theory/lifting_properties/adjunction.lean | 3 +++ src/category_theory/lifting_properties/basic.lean | 3 +++ src/category_theory/limits/shapes/strong_epi.lean | 3 +++ src/category_theory/skeletal.lean | 3 +++ src/combinatorics/additive/pluennecke_ruzsa.lean | 3 +++ src/data/qpf/multivariate/constructions/sigma.lean | 3 +++ src/field_theory/subfield.lean | 3 +++ src/group_theory/order_of_element.lean | 3 +++ src/linear_algebra/projection.lean | 3 +++ src/set_theory/ordinal/principal.lean | 3 +++ src/topology/algebra/order/group.lean | 3 +++ src/topology/continuous_function/ordered.lean | 3 +++ 18 files changed, 54 insertions(+) diff --git a/src/category_theory/adjunction/fully_faithful.lean b/src/category_theory/adjunction/fully_faithful.lean index ca29238caf89b..863d933a046ee 100644 --- a/src/category_theory/adjunction/fully_faithful.lean +++ b/src/category_theory/adjunction/fully_faithful.lean @@ -10,6 +10,9 @@ import category_theory.yoneda /-! # Adjoints of fully faithful functors +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A left adjoint is fully faithful, if and only if the unit is an isomorphism (and similarly for right adjoints and the counit). diff --git a/src/category_theory/adjunction/whiskering.lean b/src/category_theory/adjunction/whiskering.lean index 8238ecbf9a50e..5a82ed48a9f7d 100644 --- a/src/category_theory/adjunction/whiskering.lean +++ b/src/category_theory/adjunction/whiskering.lean @@ -7,6 +7,9 @@ import category_theory.whiskering import category_theory.adjunction.basic /-! +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given categories `C D E`, functors `F : D ⥤ E` and `G : E ⥤ D` with an adjunction `F ⊣ G`, we provide the induced adjunction between the functor categories `C ⥤ D` and `C ⥤ E`, diff --git a/src/category_theory/category/Cat.lean b/src/category_theory/category/Cat.lean index 3389fc80b7cac..b8ff7a982d526 100644 --- a/src/category_theory/category/Cat.lean +++ b/src/category_theory/category/Cat.lean @@ -11,6 +11,9 @@ import category_theory.bicategory.strict /-! # Category of categories +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains the definition of the category `Cat` of all categories. In this category objects are categories and morphisms are functors between these categories. diff --git a/src/category_theory/category/galois_connection.lean b/src/category_theory/category/galois_connection.lean index 2c369976fe89e..35ab915fc0bfe 100644 --- a/src/category_theory/category/galois_connection.lean +++ b/src/category_theory/category/galois_connection.lean @@ -11,6 +11,9 @@ import order.galois_connection # Galois connections between preorders are adjunctions. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + * `galois_connection.adjunction` is the adjunction associated to a galois connection. -/ diff --git a/src/category_theory/fin_category.lean b/src/category_theory/fin_category.lean index 11e88b62157be..0681cdda833d7 100644 --- a/src/category_theory/fin_category.lean +++ b/src/category_theory/fin_category.lean @@ -11,6 +11,9 @@ import category_theory.category.ulift /-! # Finite categories +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A category is finite in this sense if it has finitely many objects, and finitely many morphisms. ## Implementation diff --git a/src/category_theory/isomorphism_classes.lean b/src/category_theory/isomorphism_classes.lean index 144652c36cc4c..d58e5f1035a19 100644 --- a/src/category_theory/isomorphism_classes.lean +++ b/src/category_theory/isomorphism_classes.lean @@ -10,6 +10,9 @@ import category_theory.types /-! # Objects of a category up to an isomorphism +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + `is_isomorphic X Y := nonempty (X ≅ Y)` is an equivalence relation on the objects of a category. The quotient with respect to this relation defines a functor from our category to `Type`. -/ diff --git a/src/category_theory/lifting_properties/adjunction.lean b/src/category_theory/lifting_properties/adjunction.lean index c285c0e87de16..5602abec58826 100644 --- a/src/category_theory/lifting_properties/adjunction.lean +++ b/src/category_theory/lifting_properties/adjunction.lean @@ -11,6 +11,9 @@ import category_theory.adjunction.basic # Lifting properties and adjunction +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we obtain `adjunction.has_lifting_property_iff`, which states that when we have an adjunction `adj : G ⊣ F` between two functors `G : C ⥤ D` and `F : D ⥤ C`, then a morphism of the form `G.map i` has the left lifting diff --git a/src/category_theory/lifting_properties/basic.lean b/src/category_theory/lifting_properties/basic.lean index 79a4919486576..7c74c6abee157 100644 --- a/src/category_theory/lifting_properties/basic.lean +++ b/src/category_theory/lifting_properties/basic.lean @@ -8,6 +8,9 @@ import category_theory.comm_sq /-! # Lifting properties +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the lifting property of two morphisms in a category and shows basic properties of this notion. diff --git a/src/category_theory/limits/shapes/strong_epi.lean b/src/category_theory/limits/shapes/strong_epi.lean index ce67158f4b03e..01cfc4c59fcf4 100644 --- a/src/category_theory/limits/shapes/strong_epi.lean +++ b/src/category_theory/limits/shapes/strong_epi.lean @@ -9,6 +9,9 @@ import category_theory.lifting_properties.basic /-! # Strong epimorphisms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we define strong epimorphisms. A strong epimorphism is an epimorphism `f` which has the (unique) left lifting property with respect to monomorphisms. Similarly, a strong monomorphisms in a monomorphism which has the (unique) right lifting property diff --git a/src/category_theory/skeletal.lean b/src/category_theory/skeletal.lean index 2992995d7999e..7411364ac96e5 100644 --- a/src/category_theory/skeletal.lean +++ b/src/category_theory/skeletal.lean @@ -11,6 +11,9 @@ import category_theory.thin /-! # Skeleton of a category +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Define skeletal categories as categories in which any two isomorphic objects are equal. Construct the skeleton of an arbitrary category by taking isomorphism classes, and show it is a diff --git a/src/combinatorics/additive/pluennecke_ruzsa.lean b/src/combinatorics/additive/pluennecke_ruzsa.lean index 16eb9cd950f8d..737dad50dfabf 100644 --- a/src/combinatorics/additive/pluennecke_ruzsa.lean +++ b/src/combinatorics/additive/pluennecke_ruzsa.lean @@ -10,6 +10,9 @@ import data.rat.nnrat /-! # The Plünnecke-Ruzsa inequality +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves Ruzsa's triangle inequality, the Plünnecke-Petridis lemma, and the Plünnecke-Ruzsa inequality. diff --git a/src/data/qpf/multivariate/constructions/sigma.lean b/src/data/qpf/multivariate/constructions/sigma.lean index 784910c7c9540..462edb0482310 100644 --- a/src/data/qpf/multivariate/constructions/sigma.lean +++ b/src/data/qpf/multivariate/constructions/sigma.lean @@ -9,6 +9,9 @@ import data.qpf.multivariate.basic /-! # Dependent product and sum of QPFs are QPFs + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ universes u diff --git a/src/field_theory/subfield.lean b/src/field_theory/subfield.lean index 7d28db29ba330..dc67432a1eeff 100644 --- a/src/field_theory/subfield.lean +++ b/src/field_theory/subfield.lean @@ -10,6 +10,9 @@ import algebra.order.field.inj_surj /-! # Subfields +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Let `K` be a field. This file defines the "bundled" subfield type `subfield K`, a type whose terms correspond to subfields of `K`. This is the preferred way to talk about subfields in mathlib. Unbundled subfields (`s : set K` and `is_subfield s`) diff --git a/src/group_theory/order_of_element.lean b/src/group_theory/order_of_element.lean index 5c629adbebdea..44eb664b3460a 100644 --- a/src/group_theory/order_of_element.lean +++ b/src/group_theory/order_of_element.lean @@ -13,6 +13,9 @@ import group_theory.index /-! # Order of an element +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the order of an element of a finite group. For a finite group `G` the order of `x ∈ G` is the minimal `n ≥ 1` such that `x ^ n = 1`. diff --git a/src/linear_algebra/projection.lean b/src/linear_algebra/projection.lean index 0607c59f7ad11..d8f56f520278e 100644 --- a/src/linear_algebra/projection.lean +++ b/src/linear_algebra/projection.lean @@ -9,6 +9,9 @@ import linear_algebra.prod /-! # Projection to a subspace +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define * `linear_proj_of_is_compl (p q : submodule R E) (h : is_compl p q)`: the projection of a module `E` to a submodule `p` along its complement `q`; it is the unique linear map `f : E → p` such that diff --git a/src/set_theory/ordinal/principal.lean b/src/set_theory/ordinal/principal.lean index a79c584716972..61774bc56030a 100644 --- a/src/set_theory/ordinal/principal.lean +++ b/src/set_theory/ordinal/principal.lean @@ -9,6 +9,9 @@ import set_theory.ordinal.fixed_point /-! ### Principal ordinals +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define principal or indecomposable ordinals, and we prove the standard properties about them. ### Main definitions and results diff --git a/src/topology/algebra/order/group.lean b/src/topology/algebra/order/group.lean index a48d17fa14791..c4aac1da2e3f1 100644 --- a/src/topology/algebra/order/group.lean +++ b/src/topology/algebra/order/group.lean @@ -9,6 +9,9 @@ import topology.algebra.group.basic /-! # Topology on a linear ordered additive commutative group +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove that a linear ordered additive commutative group with order topology is a topological group. We also prove continuity of `abs : G → G` and provide convenience lemmas like `continuous_at.abs`. diff --git a/src/topology/continuous_function/ordered.lean b/src/topology/continuous_function/ordered.lean index e7073aa0bf720..7bbf9121a9ebe 100644 --- a/src/topology/continuous_function/ordered.lean +++ b/src/topology/continuous_function/ordered.lean @@ -11,6 +11,9 @@ import topology.continuous_function.basic /-! # Bundled continuous maps into orders, with order-compatible topology +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + -/ variables {α : Type*} {β : Type*} {γ : Type*} From 296726c93ccf37fa3bb8fa8c3a35e6e456ef429b Mon Sep 17 00:00:00 2001 From: leanprover-community-bot Date: Thu, 23 Feb 2023 05:34:11 +0000 Subject: [PATCH 0511/1291] chore(scripts): update nolints.txt (#18486) I am happy to remove some nolints for you! --- scripts/style-exceptions.txt | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/scripts/style-exceptions.txt b/scripts/style-exceptions.txt index 75be83d955367..3690c941e9416 100644 --- a/scripts/style-exceptions.txt +++ b/scripts/style-exceptions.txt @@ -5,7 +5,7 @@ src/control/traversable/derive.lean : line 11 : ERR_MOD : Module docstring missi src/data/array/lemmas.lean : line 9 : ERR_MOD : Module docstring missing, or too late src/data/bitvec/basic.lean : line 11 : ERR_MOD : Module docstring missing, or too late src/data/buffer/basic.lean : line 11 : ERR_MOD : Module docstring missing, or too late -src/data/qpf/multivariate/basic.lean : line 73 : ERR_LIN : Line has more than 100 characters +src/data/qpf/multivariate/basic.lean : line 76 : ERR_LIN : Line has more than 100 characters src/data/qpf/univariate/basic.lean : line 35 : ERR_LIN : Line has more than 100 characters src/data/rbmap/basic.lean : line 8 : ERR_MOD : Module docstring missing, or too late src/data/rbmap/default.lean : line 9 : ERR_MOD : Module docstring missing, or too late From 8a318021995877a44630c898d0b2bc376fceef3b Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Thu, 23 Feb 2023 10:01:28 +0000 Subject: [PATCH 0512/1291] chore(category_theory/comma): removed obviously tactic for data fields (#18440) For some `comma` categories like `structured_arrow`, `over`, `under`, the definition of some data fields could be obtained automatically by the obviously tactic. This PR removes this behaviour. Instead, specialized constructors like `structured_arrow.mk` or `structured_arrow.hom_mk` are used more systematically. --- counterexamples/pseudoelement.lean | 10 ++-------- src/category_theory/adjunction/comma.lean | 2 +- src/category_theory/comma.lean | 8 ++++---- src/category_theory/elements.lean | 2 +- src/category_theory/limits/final.lean | 6 +++--- src/category_theory/over.lean | 12 +++++------- src/category_theory/structured_arrow.lean | 14 +++++++------- src/topology/sheaves/presheaf.lean | 8 ++++---- .../sheaf_condition/pairwise_intersections.lean | 2 +- 9 files changed, 28 insertions(+), 36 deletions(-) diff --git a/counterexamples/pseudoelement.lean b/counterexamples/pseudoelement.lean index ef4f62ba6cb7d..5a28cf838eabb 100644 --- a/counterexamples/pseudoelement.lean +++ b/counterexamples/pseudoelement.lean @@ -37,17 +37,11 @@ namespace category_theory.abelian.pseudoelement /-- `x` is given by `t ↦ (t, 2 * t)`. -/ def x : over ((of ℤ ℚ) ⊞ (of ℤ ℚ)) := -begin - constructor, - exact biprod.lift (of_hom id) (of_hom (2 * id)), -end +over.mk (biprod.lift (of_hom id) (of_hom (2 * id))) /-- `y` is given by `t ↦ (t, t)`. -/ def y : over ((of ℤ ℚ) ⊞ (of ℤ ℚ)) := -begin - constructor, - exact biprod.lift (of_hom id) (of_hom id), -end +over.mk (biprod.lift (of_hom id) (of_hom id)) /-- `biprod.fst ≫ x` is pseudoequal to `biprod.fst y`. -/ lemma fst_x_pseudo_eq_fst_y : pseudo_equal _ (app biprod.fst x) (app biprod.fst y) := diff --git a/src/category_theory/adjunction/comma.lean b/src/category_theory/adjunction/comma.lean index 5b5bf955cd7fa..035c60a566ff2 100644 --- a/src/category_theory/adjunction/comma.lean +++ b/src/category_theory/adjunction/comma.lean @@ -53,7 +53,7 @@ def left_adjoint_of_structured_arrow_initials_aux (A : C) (B : D) : end, right_inv := λ f, begin - let B' : structured_arrow A G := { right := B, hom := f }, + let B' : structured_arrow A G := structured_arrow.mk f, apply (comma_morphism.w (initial.to B')).symm.trans (category.id_comp _), end } diff --git a/src/category_theory/comma.lean b/src/category_theory/comma.lean index dd632fe0d48bd..c78fa975d0b0a 100644 --- a/src/category_theory/comma.lean +++ b/src/category_theory/comma.lean @@ -58,8 +58,8 @@ variables {T : Type u₃} [category.{v₃} T] /-- The objects of the comma category are triples of an object `left : A`, an object `right : B` and a morphism `hom : L.obj left ⟶ R.obj right`. -/ structure comma (L : A ⥤ T) (R : B ⥤ T) : Type (max u₁ u₂ v₃) := -(left : A . obviously) -(right : B . obviously) +(left : A) +(right : B) (hom : L.obj left ⟶ R.obj right) -- Satisfying the inhabited linter @@ -75,8 +75,8 @@ variables {L : A ⥤ T} {R : B ⥤ T} morphisms coming from the two objects using morphisms in the image of the functors `L` and `R`. -/ @[ext] structure comma_morphism (X Y : comma L R) := -(left : X.left ⟶ Y.left . obviously) -(right : X.right ⟶ Y.right . obviously) +(left : X.left ⟶ Y.left) +(right : X.right ⟶ Y.right) (w' : L.map left ≫ Y.hom = X.hom ≫ R.map right . obviously) -- Satisfying the inhabited linter diff --git a/src/category_theory/elements.lean b/src/category_theory/elements.lean index 2229113de827e..cfe71eb4176be 100644 --- a/src/category_theory/elements.lean +++ b/src/category_theory/elements.lean @@ -120,7 +120,7 @@ def structured_arrow_equivalence : F.elements ≌ structured_arrow punit F := equivalence.mk (to_structured_arrow F) (from_structured_arrow F) (nat_iso.of_components (λ X, eq_to_iso (by tidy)) (by tidy)) (nat_iso.of_components - (λ X, { hom := { right := 𝟙 _ }, inv := { right := 𝟙 _ } }) + (λ X, structured_arrow.iso_mk (iso.refl _) (by tidy)) (by tidy)) open opposite diff --git a/src/category_theory/limits/final.lean b/src/category_theory/limits/final.lean index 4ef8f7ff0158f..18698aca45cd5 100644 --- a/src/category_theory/limits/final.lean +++ b/src/category_theory/limits/final.lean @@ -173,7 +173,7 @@ def induction {d : D} (Z : Π (X : C) (k : d ⟶ F.obj X), Sort*) begin apply nonempty.some, apply @is_preconnected_induction _ _ _ - (λ (Y : structured_arrow d F), Z Y.right Y.hom) _ _ { right := X₀, hom := k₀, } z, + (λ (Y : structured_arrow d F), Z Y.right Y.hom) _ _ (structured_arrow.mk k₀) z, { intros j₁ j₂ f a, fapply h₁ _ _ _ _ f.right _ a, convert f.w.symm, dsimp, simp, }, { intros j₁ j₂ f a, fapply h₂ _ _ _ _ f.right _ a, convert f.w.symm, dsimp, simp, }, end @@ -351,7 +351,7 @@ begin fconstructor, swap 2, fconstructor, left, fsplit, - exact { right := f, } }, + exact structured_arrow.hom_mk f (by tidy), }, case eqv_gen.refl { fconstructor, }, case eqv_gen.symm : x y h ih @@ -426,7 +426,7 @@ def induction {d : D} (Z : Π (X : C) (k : F.obj X ⟶ d), Sort*) begin apply nonempty.some, apply @is_preconnected_induction _ _ _ - (λ Y : costructured_arrow F d, Z Y.left Y.hom) _ _ { left := X₀, hom := k₀ } z, + (λ Y : costructured_arrow F d, Z Y.left Y.hom) _ _ (costructured_arrow.mk k₀) z, { intros j₁ j₂ f a, fapply h₁ _ _ _ _ f.left _ a, convert f.w, dsimp, simp, }, { intros j₁ j₂ f a, fapply h₂ _ _ _ _ f.left _ a, convert f.w, dsimp, simp, }, end diff --git a/src/category_theory/over.lean b/src/category_theory/over.lean index 032b2f16935f8..fbba8bcb73689 100644 --- a/src/category_theory/over.lean +++ b/src/category_theory/over.lean @@ -40,6 +40,7 @@ def over (X : T) := costructured_arrow (𝟭 T) X instance over.inhabited [inhabited T] : inhabited (over (default : T)) := { default := { left := default, + right := default, hom := 𝟙 _ } } namespace over @@ -222,9 +223,7 @@ variables {D : Type u₂} [category.{v₂} D] @[simps] def post (F : T ⥤ D) : over X ⥤ over (F.obj X) := { obj := λ Y, mk $ F.map Y.hom, - map := λ Y₁ Y₂ f, - { left := F.map f.left, - w' := by tidy; erw [← F.map_comp, w] } } + map := λ Y₁ Y₂ f, over.hom_mk (F.map f.left) (by tidy; erw [← F.map_comp, w]) } end @@ -238,7 +237,8 @@ def under (X : T) := structured_arrow X (𝟭 T) -- Satisfying the inhabited linter instance under.inhabited [inhabited T] : inhabited (under (default : T)) := { default := - { right := default, + { left := default, + right := default, hom := 𝟙 _ } } namespace under @@ -369,9 +369,7 @@ variables {D : Type u₂} [category.{v₂} D] @[simps] def post {X : T} (F : T ⥤ D) : under X ⥤ under (F.obj X) := { obj := λ Y, mk $ F.map Y.hom, - map := λ Y₁ Y₂ f, - { right := F.map f.right, - w' := by tidy; erw [← F.map_comp, w] } } + map := λ Y₁ Y₂ f, under.hom_mk (F.map f.right) (by tidy; erw [← F.map_comp, w]), } end diff --git a/src/category_theory/structured_arrow.lean b/src/category_theory/structured_arrow.lean index 2b289b795b53b..b977135c73d57 100644 --- a/src/category_theory/structured_arrow.lean +++ b/src/category_theory/structured_arrow.lean @@ -69,7 +69,7 @@ structured arrow given by `(X ⟶ F(U)) ⟶ (X ⟶ F(U) ⟶ F(Y))`. -/ def hom_mk' {F : C ⥤ D} {X : D} {Y : C} (U : structured_arrow X F) (f : U.right ⟶ Y) : -U ⟶ mk (U.hom ≫ F.map f) := { right := f } +U ⟶ mk (U.hom ≫ F.map f) := { left := eq_to_hom (by ext), right := f } /-- To construct an isomorphism of structured arrows, @@ -165,9 +165,9 @@ comma.pre_right _ F G /-- The functor `(S, F) ⥤ (G(S), F ⋙ G)`. -/ @[simps] def post (S : C) (F : B ⥤ C) (G : C ⥤ D) : structured_arrow S F ⥤ structured_arrow (G.obj S) (F ⋙ G) := -{ obj := λ X, { right := X.right, hom := G.map X.hom }, - map := λ X Y f, { right := f.right, w' := - by { simp [functor.comp_map, ←G.map_comp, ← f.w] } } } +{ obj := λ X, structured_arrow.mk (G.map X.hom), + map := λ X Y f, structured_arrow.hom_mk f.right + (by simp [functor.comp_map, ←G.map_comp, ← f.w]) } instance small_proj_preimage_of_locally_small {𝒢 : set C} [small.{v₁} 𝒢] [locally_small.{v₁} D] : small.{v₁} ((proj S T).obj ⁻¹' 𝒢) := @@ -315,9 +315,9 @@ comma.pre_left F G _ /-- The functor `(F, S) ⥤ (F ⋙ G, G(S))`. -/ @[simps] def post (F : B ⥤ C) (G : C ⥤ D) (S : C) : costructured_arrow F S ⥤ costructured_arrow (F ⋙ G) (G.obj S) := -{ obj := λ X, { left := X.left, hom := G.map X.hom }, - map := λ X Y f, { left := f.left, w' := - by { simp [functor.comp_map, ←G.map_comp, ← f.w] } } } +{ obj := λ X, costructured_arrow.mk (G.map X.hom), + map := λ X Y f, costructured_arrow.hom_mk f.left + (by simp [functor.comp_map, ←G.map_comp, ← f.w]), } instance small_proj_preimage_of_locally_small {𝒢 : set C} [small.{v₁} 𝒢] [locally_small.{v₁} D] : small.{v₁} ((proj S T).obj ⁻¹' 𝒢) := diff --git a/src/topology/sheaves/presheaf.lean b/src/topology/sheaves/presheaf.lean index 70726dfecf1cc..6e63c82002e36 100644 --- a/src/topology/sheaves/presheaf.lean +++ b/src/topology/sheaves/presheaf.lean @@ -241,10 +241,10 @@ def pullback_map {X Y : Top.{v}} (f : X ⟶ Y) {ℱ 𝒢 : Y.presheaf C} (α : def pullback_obj_obj_of_image_open {X Y : Top.{v}} (f : X ⟶ Y) (ℱ : Y.presheaf C) (U : opens X) (H : is_open (f '' U)) : (pullback_obj f ℱ).obj (op U) ≅ ℱ.obj (op ⟨_, H⟩) := begin - let x : costructured_arrow (opens.map f).op (op U) := - { left := op ⟨f '' U, H⟩, - hom := ((@hom_of_le _ _ _ ((opens.map f).obj ⟨_, H⟩) (set.image_preimage.le_u_l _)).op : - op ((opens.map f).obj (⟨⇑f '' ↑U, H⟩)) ⟶ op U) }, + let x : costructured_arrow (opens.map f).op (op U) := begin + refine @costructured_arrow.mk _ _ _ _ _ (op (opens.mk (f '' U) H)) _ _, + exact ((@hom_of_le _ _ _ ((opens.map f).obj ⟨_, H⟩) (set.image_preimage.le_u_l _)).op), + end, have hx : is_terminal x := { lift := λ s, begin diff --git a/src/topology/sheaves/sheaf_condition/pairwise_intersections.lean b/src/topology/sheaves/sheaf_condition/pairwise_intersections.lean index 56368c81bbc3f..7ad353b49cde6 100644 --- a/src/topology/sheaves/sheaf_condition/pairwise_intersections.lean +++ b/src/topology/sheaves/sheaf_condition/pairwise_intersections.lean @@ -113,7 +113,7 @@ def pairwise_to_opens_le_cover : pairwise ι ⥤ opens_le_cover U := instance (V : opens_le_cover U) : nonempty (structured_arrow V (pairwise_to_opens_le_cover U)) := -⟨{ right := single (V.index), hom := V.hom_to_index }⟩ +⟨@structured_arrow.mk _ _ _ _ _ (single (V.index)) _ (by exact V.hom_to_index)⟩ /-- The diagram consisting of the `U i` and `U i ⊓ U j` is cofinal in the diagram From 1ed3a113dbc6f5b33eae3b96211d4e26ca3a5e9d Mon Sep 17 00:00:00 2001 From: Aaron Anderson Date: Thu, 23 Feb 2023 11:46:48 +0000 Subject: [PATCH 0513/1291] feat(model_theory/order): Renames and generalizes notation about ordered structures (#17870) Generalizes the definition of first-order sentences and theories about orders to work in ordered languages rather than just the one-symbol language `language.order`. Renames these sentences and theories accordingly: for instance, the dot notation `L.preorder_theory` is the theory of preordered `L`-structures, which requires that `preorder_theory` be directly in the namespace `first_order.language`. --- src/model_theory/order.lean | 94 ++++++++++++++++++++----------------- 1 file changed, 51 insertions(+), 43 deletions(-) diff --git a/src/model_theory/order.lean b/src/model_theory/order.lean index 640ce8e9546ed..c14ca9cadf61d 100644 --- a/src/model_theory/order.lean +++ b/src/model_theory/order.lean @@ -14,13 +14,13 @@ This file defines ordered first-order languages and structures, as well as their * `first_order.language.order_Structure` is the structure on an ordered type, assigning the symbol representing `≤` to the actual relation `≤`. * `first_order.language.is_ordered` points out a specific symbol in a language as representing `≤`. -* `first_order.language.is_ordered_structure` indicates that a structure over a -* `first_order.language.Theory.linear_order` and similar define the theories of preorders, +* `first_order.language.ordered_structure` indicates that the `≤` symbol in an ordered language +is interpreted as the actual relation `≤` in a particular structure. +* `first_order.language.linear_order_theory` and similar define the theories of preorders, partial orders, and linear orders. -* `first_order.language.Theory.DLO` defines the theory of dense linear orders without endpoints, a +* `first_order.language.DLO` defines the theory of dense linear orders without endpoints, a particularly useful example in model theory. - ## Main Results * `partial_order`s model the theory of partial orders, `linear_order`s model the theory of linear orders, and dense linear orders without endpoints model `Theory.DLO`. @@ -40,11 +40,11 @@ variables {L : language.{u v}} {α : Type w} {M : Type w'} {n : ℕ} protected def order : language := language.mk₂ empty empty empty empty unit -namespace order - -instance Structure [has_le M] : language.order.Structure M := +instance order_Structure [has_le M] : language.order.Structure M := Structure.mk₂ empty.elim empty.elim empty.elim empty.elim (λ _, (≤)) +namespace order + instance : is_relational (language.order) := language.is_relational_mk₂ instance : subsingleton (language.order.relations n) := @@ -89,108 +89,114 @@ Lhom.funext (subsingleton.elim _ _) (subsingleton.elim _ _) instance : is_ordered (L.sum language.order) := ⟨sum.inr is_ordered.le_symb⟩ +section +variables (L) [is_ordered L] + /-- The theory of preorders. -/ -protected def Theory.preorder : language.order.Theory := +def preorder_theory : L.Theory := {le_symb.reflexive, le_symb.transitive} /-- The theory of partial orders. -/ -protected def Theory.partial_order : language.order.Theory := +def partial_order_theory : L.Theory := {le_symb.reflexive, le_symb.antisymmetric, le_symb.transitive} /-- The theory of linear orders. -/ -protected def Theory.linear_order : language.order.Theory := +def linear_order_theory : L.Theory := {le_symb.reflexive, le_symb.antisymmetric, le_symb.transitive, le_symb.total} /-- A sentence indicating that an order has no top element: $\forall x, \exists y, \neg y \le x$. -/ -protected def sentence.no_top_order : language.order.sentence := ∀' ∃' ∼ ((&1).le &0) +def no_top_order_sentence : L.sentence := ∀' ∃' ∼ ((&1).le &0) /-- A sentence indicating that an order has no bottom element: $\forall x, \exists y, \neg x \le y$. -/ -protected def sentence.no_bot_order : language.order.sentence := ∀' ∃' ∼ ((&0).le &1) +def no_bot_order_sentence : L.sentence := ∀' ∃' ∼ ((&0).le &1) /-- A sentence indicating that an order is dense: $\forall x, \forall y, x < y \to \exists z, x < z \wedge z < y$. -/ -protected def sentence.densely_ordered : language.order.sentence := +def densely_ordered_sentence : L.sentence := ∀' ∀' (((&0).lt &1) ⟹ (∃' (((&0).lt &2) ⊓ ((&2).lt &1)))) /-- The theory of dense linear orders without endpoints. -/ -protected def Theory.DLO : language.order.Theory := -Theory.linear_order ∪ {sentence.no_top_order, sentence.no_bot_order, sentence.densely_ordered} +def DLO : L.Theory := +L.linear_order_theory ∪ + {L.no_top_order_sentence, L.no_bot_order_sentence, L.densely_ordered_sentence} + +end variables (L M) /-- A structure is ordered if its language has a `≤` symbol whose interpretation is -/ -abbreviation is_ordered_structure [is_ordered L] [has_le M] [L.Structure M] : Prop := +abbreviation ordered_structure [is_ordered L] [has_le M] [L.Structure M] : Prop := Lhom.is_expansion_on (order_Lhom L) M variables {L M} -@[simp] lemma is_ordered_structure_iff [is_ordered L] [has_le M] [L.Structure M] : - L.is_ordered_structure M ↔ Lhom.is_expansion_on (order_Lhom L) M := iff.rfl +@[simp] lemma ordered_structure_iff [is_ordered L] [has_le M] [L.Structure M] : + L.ordered_structure M ↔ Lhom.is_expansion_on (order_Lhom L) M := iff.rfl -instance is_ordered_structure_has_le [has_le M] : - is_ordered_structure language.order M := +instance ordered_structure_has_le [has_le M] : + ordered_structure language.order M := begin - rw [is_ordered_structure_iff, order_Lhom_order], + rw [ordered_structure_iff, order_Lhom_order], exact Lhom.id_is_expansion_on M, end instance model_preorder [preorder M] : - M ⊨ Theory.preorder := + M ⊨ language.order.preorder_theory := begin - simp only [Theory.preorder, Theory.model_iff, set.mem_insert_iff, set.mem_singleton_iff, + simp only [preorder_theory, Theory.model_iff, set.mem_insert_iff, set.mem_singleton_iff, forall_eq_or_imp, relations.realize_reflexive, rel_map_apply₂, forall_eq, relations.realize_transitive], exact ⟨le_refl, λ _ _ _, le_trans⟩ end instance model_partial_order [partial_order M] : - M ⊨ Theory.partial_order := + M ⊨ language.order.partial_order_theory := begin - simp only [Theory.partial_order, Theory.model_iff, set.mem_insert_iff, set.mem_singleton_iff, + simp only [partial_order_theory, Theory.model_iff, set.mem_insert_iff, set.mem_singleton_iff, forall_eq_or_imp, relations.realize_reflexive, rel_map_apply₂, relations.realize_antisymmetric, forall_eq, relations.realize_transitive], exact ⟨le_refl, λ _ _, le_antisymm, λ _ _ _, le_trans⟩, end instance model_linear_order [linear_order M] : - M ⊨ Theory.linear_order := + M ⊨ language.order.linear_order_theory := begin - simp only [Theory.linear_order, Theory.model_iff, set.mem_insert_iff, set.mem_singleton_iff, + simp only [linear_order_theory, Theory.model_iff, set.mem_insert_iff, set.mem_singleton_iff, forall_eq_or_imp, relations.realize_reflexive, rel_map_apply₂, relations.realize_antisymmetric, relations.realize_transitive, forall_eq, relations.realize_total], exact ⟨le_refl, λ _ _, le_antisymm, λ _ _ _, le_trans, le_total⟩, end -section is_ordered_structure +section ordered_structure variables [is_ordered L] [L.Structure M] -@[simp] lemma rel_map_le_symb [has_le M] [L.is_ordered_structure M] {a b : M} : +@[simp] lemma rel_map_le_symb [has_le M] [L.ordered_structure M] {a b : M} : rel_map (le_symb : L.relations 2) ![a, b] ↔ a ≤ b := begin rw [← order_Lhom_le_symb, Lhom.map_on_relation], refl, end -@[simp] lemma term.realize_le [has_le M] [L.is_ordered_structure M] +@[simp] lemma term.realize_le [has_le M] [L.ordered_structure M] {t₁ t₂ : L.term (α ⊕ fin n)} {v : α → M} {xs : fin n → M} : (t₁.le t₂).realize v xs ↔ t₁.realize (sum.elim v xs) ≤ t₂.realize (sum.elim v xs) := by simp [term.le] -@[simp] lemma term.realize_lt [preorder M] [L.is_ordered_structure M] +@[simp] lemma term.realize_lt [preorder M] [L.ordered_structure M] {t₁ t₂ : L.term (α ⊕ fin n)} {v : α → M} {xs : fin n → M} : (t₁.lt t₂).realize v xs ↔ t₁.realize (sum.elim v xs) < t₂.realize (sum.elim v xs) := by simp [term.lt, lt_iff_le_not_le] -end is_ordered_structure +end ordered_structure section has_le variables [has_le M] -theorem realize_no_top_order_iff : M ⊨ sentence.no_top_order ↔ no_top_order M := +theorem realize_no_top_order_iff : M ⊨ language.order.no_top_order_sentence ↔ no_top_order M := begin - simp only [sentence.no_top_order, sentence.realize, formula.realize, bounded_formula.realize_all, + simp only [no_top_order_sentence, sentence.realize, formula.realize, bounded_formula.realize_all, bounded_formula.realize_ex, bounded_formula.realize_not, realize, term.realize_le, sum.elim_inr], refine ⟨λ h, ⟨λ a, h a⟩, _⟩, @@ -198,12 +204,13 @@ begin exact exists_not_le a, end -@[simp] lemma realize_no_top_order [h : no_top_order M] : M ⊨ sentence.no_top_order := +@[simp] lemma realize_no_top_order [h : no_top_order M] : + M ⊨ language.order.no_top_order_sentence := realize_no_top_order_iff.2 h -theorem realize_no_bot_order_iff : M ⊨ sentence.no_bot_order ↔ no_bot_order M := +theorem realize_no_bot_order_iff : M ⊨ language.order.no_bot_order_sentence ↔ no_bot_order M := begin - simp only [sentence.no_bot_order, sentence.realize, formula.realize, bounded_formula.realize_all, + simp only [no_bot_order_sentence, sentence.realize, formula.realize, bounded_formula.realize_all, bounded_formula.realize_ex, bounded_formula.realize_not, realize, term.realize_le, sum.elim_inr], refine ⟨λ h, ⟨λ a, h a⟩, _⟩, @@ -211,15 +218,16 @@ begin exact exists_not_ge a, end -@[simp] lemma realize_no_bot_order [h : no_bot_order M] : M ⊨ sentence.no_bot_order := +@[simp] lemma realize_no_bot_order [h : no_bot_order M] : + M ⊨ language.order.no_bot_order_sentence := realize_no_bot_order_iff.2 h end has_le theorem realize_densely_ordered_iff [preorder M] : - M ⊨ sentence.densely_ordered ↔ densely_ordered M := + M ⊨ language.order.densely_ordered_sentence ↔ densely_ordered M := begin - simp only [sentence.densely_ordered, sentence.realize, formula.realize, + simp only [densely_ordered_sentence, sentence.realize, formula.realize, bounded_formula.realize_imp, bounded_formula.realize_all, realize, term.realize_lt, sum.elim_inr, bounded_formula.realize_ex, bounded_formula.realize_inf], refine ⟨λ h, ⟨λ a b ab, h a b ab⟩, _⟩, @@ -228,13 +236,13 @@ begin end @[simp] lemma realize_densely_ordered [preorder M] [h : densely_ordered M] : - M ⊨ sentence.densely_ordered := + M ⊨ language.order.densely_ordered_sentence := realize_densely_ordered_iff.2 h instance model_DLO [linear_order M] [densely_ordered M] [no_top_order M] [no_bot_order M] : - M ⊨ Theory.DLO := + M ⊨ language.order.DLO := begin - simp only [Theory.DLO, set.union_insert, set.union_singleton, Theory.model_iff, + simp only [DLO, set.union_insert, set.union_singleton, Theory.model_iff, set.mem_insert_iff, forall_eq_or_imp, realize_no_top_order, realize_no_bot_order, realize_densely_ordered, true_and], rw ← Theory.model_iff, From b8d2eaa69d69ce8f03179a5cda774fc0cde984e4 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 23 Feb 2023 11:46:49 +0000 Subject: [PATCH 0514/1291] refactor(algebra/{dual_number,triv_sq_zero_ext}): support non-commutative rings (#18384) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This changes the definition of multiplication on `triv_sq_zero_ext R M` as follows ```diff -x * y = (x.1 * y.1, x.1 • y.2 + y.1 • x.2) +x * y = (x.1 * y.1, x.1 • y.2 + op y.1 • x.2) ``` which for `R=M` gives the more natural `x.1 * y.2 + x.2 * y.1` in the second term instead of `x.1 * y.2 + y.1 * x.2`. This adds some slightly annoying typeclass arguments that could eventually be cleaned up by #10716; but that PR has rotted under the mathlib4 tide. Either way, the weird typeclass arguments needed with `triv_sq_zero_ext R M` are all invisible when working with `dual_number R`. This is motivated by wanting to talk about `dual_number (quaternion R)` or `dual_number (matrix n n R)`. Unfortunately this breaks the `topological_semiring` instance (making it need an `@` and trigger the `fails_quickly` linter), but this instance isn't really interesting anyway. --- src/algebra/dual_number.lean | 2 +- src/algebra/triv_sq_zero_ext.lean | 261 ++++++++++++------ .../normed_space/triv_sq_zero_ext.lean | 21 +- .../exterior_algebra/basic.lean | 15 +- src/linear_algebra/tensor_algebra/basic.lean | 15 +- src/topology/instances/triv_sq_zero_ext.lean | 22 +- 6 files changed, 234 insertions(+), 102 deletions(-) diff --git a/src/algebra/dual_number.lean b/src/algebra/dual_number.lean index 9aa9098f64a4c..9daad87c0525c 100644 --- a/src/algebra/dual_number.lean +++ b/src/algebra/dual_number.lean @@ -56,7 +56,7 @@ open triv_sq_zero_ext @[simp] lemma snd_eps [has_zero R] [has_one R] : snd ε = (1 : R) := snd_inr _ _ /-- A version of `triv_sq_zero_ext.snd_mul` with `*` instead of `•`. -/ -@[simp] lemma snd_mul [semiring R] (x y : R[ε]) : snd (x * y) = fst x * snd y + fst y * snd x := +@[simp] lemma snd_mul [semiring R] (x y : R[ε]) : snd (x * y) = fst x * snd y + snd x * fst y := snd_mul _ _ @[simp] lemma eps_mul_eps [semiring R] : (ε * ε : R[ε]) = 0 := inr_mul_inr _ _ _ diff --git a/src/algebra/triv_sq_zero_ext.lean b/src/algebra/triv_sq_zero_ext.lean index fe789d9ff6354..c24d6f137553b 100644 --- a/src/algebra/triv_sq_zero_ext.lean +++ b/src/algebra/triv_sq_zero_ext.lean @@ -10,12 +10,27 @@ import linear_algebra.prod /-! # Trivial Square-Zero Extension -Given a module `M` over a ring `R`, the trivial square-zero extension of `M` over `R` is defined -to be the `R`-algebra `R ⊕ M` with multiplication given by -`(r₁ + m₁) * (r₂ + m₂) = r₁ r₂ + r₁ m₂ + r₂ m₁`. +Given a ring `R` together with an `(R, R)`-bimodule `M`, the trivial square-zero extension of `M` +over `R` is defined to be the `R`-algebra `R ⊕ M` with multiplication given by +`(r₁ + m₁) * (r₂ + m₂) = r₁ r₂ + r₁ m₂ + m₁ r₂`. It is a square-zero extension because `M^2 = 0`. +Note that expressing this requires bimodules; we write these in general for a +not-necessarily-commutative `R` as: +```lean +variables {R M : Type*} [semiring R] [add_comm_monoid M] +variables [module R M] [module Rᵐᵒᵖ M] [smul_comm_class R Rᵐᵒᵖ M] +``` +If we instead work with a commutative `R'` acting symmetrically on `M`, we write +```lean +variables {R' M : Type*} [comm_semiring R'] [add_comm_monoid M] +variables [module R' M] [module R'ᵐᵒᵖ M] [is_central_scalar R' M] +``` +noting that in this context `is_central_scalar R' M` implies `smul_comm_class R' R'ᵐᵒᵖ M`. + +Many of the later results in this file are only stated for the commutative `R'` for simplicity. + ## Main definitions * `triv_sq_zero_ext.inl`, `triv_sq_zero_ext.inr`: the canonical inclusions into @@ -47,6 +62,8 @@ local notation `tsze` := triv_sq_zero_ext namespace triv_sq_zero_ext +open mul_opposite (op) + section basic variables {R : Type u} {M : Type v} @@ -275,27 +292,29 @@ variables {R : Type u} {M : Type v} instance [has_one R] [has_zero M] : has_one (tsze R M) := ⟨(1, 0)⟩ -instance [has_mul R] [has_add M] [has_smul R M] : has_mul (tsze R M) := -⟨λ x y, (x.1 * y.1, x.1 • y.2 + y.1 • x.2)⟩ +instance [has_mul R] [has_add M] [has_smul R M] [has_smul Rᵐᵒᵖ M] : has_mul (tsze R M) := +⟨λ x y, (x.1 * y.1, x.1 • y.2 + op y.1 • x.2)⟩ @[simp] lemma fst_one [has_one R] [has_zero M] : (1 : tsze R M).fst = 1 := rfl @[simp] lemma snd_one [has_one R] [has_zero M] : (1 : tsze R M).snd = 0 := rfl -@[simp] lemma fst_mul [has_mul R] [has_add M] [has_smul R M] (x₁ x₂ : tsze R M) : +@[simp] lemma fst_mul [has_mul R] [has_add M] [has_smul R M] [has_smul Rᵐᵒᵖ M] (x₁ x₂ : tsze R M) : (x₁ * x₂).fst = x₁.fst * x₂.fst := rfl -@[simp] lemma snd_mul [has_mul R] [has_add M] [has_smul R M] (x₁ x₂ : tsze R M) : - (x₁ * x₂).snd = x₁.fst • x₂.snd + x₂.fst • x₁.snd := rfl +@[simp] lemma snd_mul [has_mul R] [has_add M] [has_smul R M] [has_smul Rᵐᵒᵖ M] (x₁ x₂ : tsze R M) : + (x₁ * x₂).snd = x₁.fst • x₂.snd + op x₂.fst • x₁.snd := rfl section variables (M) @[simp] lemma inl_one [has_one R] [has_zero M] : (inl 1 : tsze R M) = 1 := rfl -@[simp] lemma inl_mul [monoid R] [add_monoid M] [distrib_mul_action R M] (r₁ r₂ : R) : +@[simp] lemma inl_mul [monoid R] [add_monoid M] + [distrib_mul_action R M] [distrib_mul_action Rᵐᵒᵖ M] (r₁ r₂ : R) : (inl (r₁ * r₂) : tsze R M) = inl r₁ * inl r₂ := -ext rfl $ show (0 : M) = r₁ • 0 + r₂ • 0, by rw [smul_zero, zero_add, smul_zero] +ext rfl $ show (0 : M) = r₁ • 0 + op r₂ • 0, by rw [smul_zero, zero_add, smul_zero] -lemma inl_mul_inl [monoid R] [add_monoid M] [distrib_mul_action R M] (r₁ r₂ : R) : +lemma inl_mul_inl [monoid R] [add_monoid M] + [distrib_mul_action R M] [distrib_mul_action Rᵐᵒᵖ M] (r₁ r₂ : R) : (inl r₁ * inl r₂ : tsze R M) = inl (r₁ * r₂) := (inl_mul M r₁ r₂).symm @@ -304,24 +323,26 @@ end section variables (R) -@[simp] lemma inr_mul_inr [semiring R] [add_comm_monoid M] [module R M] (m₁ m₂ : M) : +@[simp] lemma inr_mul_inr [semiring R] [add_comm_monoid M] [module R M] [module Rᵐᵒᵖ M] + (m₁ m₂ : M) : (inr m₁ * inr m₂ : tsze R M) = 0 := -ext (mul_zero _) $ show (0 : R) • m₂ + (0 : R) • m₁ = 0, by rw [zero_smul, zero_add, zero_smul] +ext (mul_zero _) $ show (0 : R) • m₂ + (0 : Rᵐᵒᵖ) • m₁ = 0, by rw [zero_smul, zero_add, zero_smul] end -lemma inl_mul_inr [semiring R] [add_comm_monoid M] [module R M] (r : R) (m : M) : +lemma inl_mul_inr [semiring R] [add_comm_monoid M] [module R M] [module Rᵐᵒᵖ M] (r : R) (m : M) : (inl r * inr m : tsze R M) = inr (r • m) := -ext (mul_zero r) $ show r • m + (0 : R) • 0 = r • m, by rw [smul_zero, add_zero] +ext (mul_zero r) $ show r • m + (0 : Rᵐᵒᵖ) • 0 = r • m, by rw [smul_zero, add_zero] -lemma inr_mul_inl [semiring R] [add_comm_monoid M] [module R M] (r : R) (m : M) : - (inr m * inl r : tsze R M) = inr (r • m) := -ext (zero_mul r) $ show (0 : R) • 0 + r • m = r • m, by rw [smul_zero, zero_add] +lemma inr_mul_inl [semiring R] [add_comm_monoid M] [module R M] [module Rᵐᵒᵖ M] (r : R) (m : M) : + (inr m * inl r : tsze R M) = inr (op r • m) := +ext (zero_mul r) $ show (0 : R) • 0 + op r • m = op r • m, by rw [smul_zero, zero_add] -instance [monoid R] [add_monoid M] [distrib_mul_action R M] : mul_one_class (tsze R M) := -{ one_mul := λ x, ext (one_mul x.1) $ show (1 : R) • x.2 + x.1 • 0 = x.2, +instance [monoid R] [add_monoid M] [distrib_mul_action R M] [distrib_mul_action Rᵐᵒᵖ M] : + mul_one_class (tsze R M) := +{ one_mul := λ x, ext (one_mul x.1) $ show (1 : R) • x.2 + op x.1 • 0 = x.2, by rw [one_smul, smul_zero, add_zero], - mul_one := λ x, ext (mul_one x.1) $ show (x.1 • 0 : M) + (1 : R) • x.2 = x.2, + mul_one := λ x, ext (mul_one x.1) $ show (x.1 • 0 : M) + (1 : Rᵐᵒᵖ) • x.2 = x.2, by rw [smul_zero, zero_add, one_smul], .. triv_sq_zero_ext.has_one, .. triv_sq_zero_ext.has_mul } @@ -354,69 +375,139 @@ instance [add_group_with_one R] [add_group M] : add_group_with_one (tsze R M) := @[simp] lemma inl_int_cast [add_group_with_one R] [add_group M] (z : ℤ) : (inl z : tsze R M) = z := rfl -instance [semiring R] [add_comm_monoid M] [module R M] : non_assoc_semiring (tsze R M) := -{ zero_mul := λ x, ext (zero_mul x.1) $ show (0 : R) • x.2 + x.1 • 0 = 0, +instance [semiring R] [add_comm_monoid M] [module R M] [module Rᵐᵒᵖ M] : + non_assoc_semiring (tsze R M) := +{ zero_mul := λ x, ext (zero_mul x.1) $ show (0 : R) • x.2 + op x.1 • 0 = 0, by rw [zero_smul, zero_add, smul_zero], - mul_zero := λ x, ext (mul_zero x.1) $ show (x.1 • 0 : M) + (0 : R) • x.2 = 0, + mul_zero := λ x, ext (mul_zero x.1) $ show (x.1 • 0 : M) + (0 : Rᵐᵒᵖ) • x.2 = 0, by rw [smul_zero, zero_add, zero_smul], left_distrib := λ x₁ x₂ x₃, ext (mul_add x₁.1 x₂.1 x₃.1) $ - show x₁.1 • (x₂.2 + x₃.2) + (x₂.1 + x₃.1) • x₁.2 = - x₁.1 • x₂.2 + x₂.1 • x₁.2 + (x₁.1 • x₃.2 + x₃.1 • x₁.2), + show x₁.1 • (x₂.2 + x₃.2) + (op x₂.1 + op x₃.1) • x₁.2 = + x₁.1 • x₂.2 + op x₂.1 • x₁.2 + (x₁.1 • x₃.2 + op x₃.1 • x₁.2), by simp_rw [smul_add, add_smul, add_add_add_comm], right_distrib := λ x₁ x₂ x₃, ext (add_mul x₁.1 x₂.1 x₃.1) $ - show (x₁.1 + x₂.1) • x₃.2 + x₃.1 • (x₁.2 + x₂.2) = - x₁.1 • x₃.2 + x₃.1 • x₁.2 + (x₂.1 • x₃.2 + x₃.1 • x₂.2), + show (x₁.1 + x₂.1) • x₃.2 + op x₃.1 • (x₁.2 + x₂.2) = + x₁.1 • x₃.2 + op x₃.1 • x₁.2 + (x₂.1 • x₃.2 + op x₃.1 • x₂.2), by simp_rw [add_smul, smul_add, add_add_add_comm], .. triv_sq_zero_ext.add_monoid_with_one, .. triv_sq_zero_ext.mul_one_class, .. triv_sq_zero_ext.add_comm_monoid } -instance [ring R] [add_comm_group M] [module R M] : non_assoc_ring (tsze R M) := +instance [ring R] [add_comm_group M] [module R M] [module Rᵐᵒᵖ M] : + non_assoc_ring (tsze R M) := { .. triv_sq_zero_ext.add_group_with_one, .. triv_sq_zero_ext.non_assoc_semiring } -instance [comm_monoid R] [add_monoid M] [distrib_mul_action R M] : has_pow (tsze R M) ℕ := -⟨λ x n, ⟨x.fst^n, n • x.fst ^ n.pred • x.snd⟩⟩ +open_locale big_operators + +/-- In the general non-commutative case, the power operator is -@[simp] lemma fst_pow [comm_monoid R] [add_monoid M] [distrib_mul_action R M] +$$\begin{align} +(r + m)^n &= r^n + r^{n-1}m + r^{n-2}mr + \cdots + rmr^{n-2} + mr^{n-1} \\ + & =r^n + \sum_{i = 0}^{n - 1} r^{(n - 1) - i} m r^{i} +\end{align}$$ + +In the commutative case this becomes the simpler $(r + m)^n = r^n + nr^{n-1}m$. +-/ +instance [monoid R] [add_monoid M] [distrib_mul_action R M] [distrib_mul_action Rᵐᵒᵖ M] : + has_pow (tsze R M) ℕ := +⟨λ x n, ⟨x.fst^n, + ((list.range n).map (λ i, x.fst ^ (n.pred - i) • op (x.fst ^ i) • x.snd)).sum⟩⟩ + +@[simp] lemma fst_pow [monoid R] [add_monoid M] [distrib_mul_action R M] [distrib_mul_action Rᵐᵒᵖ M] (x : tsze R M) (n : ℕ) : fst (x ^ n) = x.fst ^ n := rfl -@[simp] lemma snd_pow [comm_monoid R] [add_monoid M] [distrib_mul_action R M] +lemma snd_pow_eq_sum [monoid R] [add_monoid M] + [distrib_mul_action R M] [distrib_mul_action Rᵐᵒᵖ M] (x : tsze R M) (n : ℕ) : + snd (x ^ n) = ((list.range n).map (λ i, x.fst ^ (n.pred - i) • op (x.fst ^ i) • x.snd)).sum := rfl + +lemma snd_pow_of_smul_comm [monoid R] [add_monoid M] + [distrib_mul_action R M] [distrib_mul_action Rᵐᵒᵖ M] [smul_comm_class R Rᵐᵒᵖ M] + (x : tsze R M) (n : ℕ) (h : op x.fst • x.snd = x.fst • x.snd) : + snd (x ^ n) = n • x.fst ^ n.pred • x.snd := +begin + have : ∀ n : ℕ, op (x.fst ^ n) • x.snd = x.fst ^ n • x.snd, + { intro n, + induction n with n ih, + { simp }, + { rw [pow_succ', mul_opposite.op_mul, mul_smul, mul_smul, ←h, + smul_comm (_ : R) (op x.fst) x.snd, ih] } }, + simp_rw [snd_pow_eq_sum, this, smul_smul, ←pow_add], + cases n, + { rw [nat.pred_zero, pow_zero, list.range_zero, zero_smul, list.map_nil, list.sum_nil] }, + simp_rw nat.pred_succ, + refine (list.sum_eq_card_nsmul _ (x.fst ^ n • x.snd) _).trans _, + { rintros m hm, + simp_rw [list.mem_map, list.mem_range] at hm, + obtain ⟨i, hi, rfl⟩ := hm, + rw tsub_add_cancel_of_le (nat.lt_succ_iff.mp hi) }, + { rw [list.length_map, list.length_range] } +end + +@[simp] lemma snd_pow [comm_monoid R] [add_monoid M] + [distrib_mul_action R M] [distrib_mul_action Rᵐᵒᵖ M] [is_central_scalar R M] (x : tsze R M) (n : ℕ) : - snd (x ^ n) = n • x.fst ^ n.pred • x.snd := rfl + snd (x ^ n) = n • x.fst ^ n.pred • x.snd := +snd_pow_of_smul_comm _ _ (op_smul_eq_smul _ _) -@[simp] lemma inl_pow [comm_monoid R] [add_monoid M] [distrib_mul_action R M] +@[simp] lemma inl_pow [monoid R] [add_monoid M] [distrib_mul_action R M] [distrib_mul_action Rᵐᵒᵖ M] (r : R) (n : ℕ) : (inl r ^ n : tsze R M) = inl (r ^ n) := -ext rfl $ by simp +ext rfl $ by simp [snd_pow_eq_sum] -instance [comm_monoid R] [add_monoid M] [distrib_mul_action R M] : monoid (tsze R M) := +instance [monoid R] [add_monoid M] + [distrib_mul_action R M] [distrib_mul_action Rᵐᵒᵖ M] [smul_comm_class R Rᵐᵒᵖ M] : + monoid (tsze R M) := { mul_assoc := λ x y z, ext (mul_assoc x.1 y.1 z.1) $ - show (x.1 * y.1) • z.2 + z.1 • (x.1 • y.2 + y.1 • x.2) = - x.1 • (y.1 • z.2 + z.1 • y.2) + (y.1 * z.1) • x.2, - by simp_rw [smul_add, ← mul_smul, add_assoc, mul_comm], + show (x.1 * y.1) • z.2 + op z.1 • (x.1 • y.2 + op y.1 • x.2) = + x.1 • (y.1 • z.2 + op z.1 • y.2) + (op z.1 * op y.1) • x.2, + by simp_rw [smul_add, ← mul_smul, add_assoc, smul_comm], npow := λ n x, x ^ n, - npow_zero' := λ x, ext (pow_zero x.fst) (zero_smul _ _), + npow_zero' := λ x, ext (pow_zero x.fst) (by simp [snd_pow_eq_sum]), npow_succ' := λ n x, ext (pow_succ _ _) begin - dsimp, - rw [smul_comm (_ : R) n (_ : M), smul_smul, succ_nsmul'], + simp_rw [snd_mul, snd_pow_eq_sum, nat.pred_succ], cases n, - { simp_rw [zero_smul ]}, - { rw [nat.pred_succ, pow_succ] } + { simp [list.range_succ], }, + simp_rw [nat.pred_succ], + rw [list.range_succ, list.map_append, list.sum_append, list.map_singleton, list.sum_singleton, + nat.sub_self, pow_zero, one_smul, list.smul_sum, list.map_map, function.comp, fst_pow], + simp_rw [smul_smul, ←pow_succ, nat.succ_eq_add_one], + congr' 2, + refine list.map_congr (λ i hi, _), + rw [list.mem_range, nat.lt_succ_iff] at hi, + rw nat.sub_add_comm hi, end, .. triv_sq_zero_ext.mul_one_class } -instance [comm_monoid R] [add_comm_monoid M] [distrib_mul_action R M] : comm_monoid (tsze R M) := +instance [semiring R] [add_comm_monoid M] + [module R M] [module Rᵐᵒᵖ M] [smul_comm_class R Rᵐᵒᵖ M] : + semiring (tsze R M) := +{ .. triv_sq_zero_ext.monoid, + .. triv_sq_zero_ext.non_assoc_semiring } + +instance [ring R] [add_comm_group M] + [module R M] [module Rᵐᵒᵖ M] [smul_comm_class R Rᵐᵒᵖ M] : + ring (tsze R M) := +{ .. triv_sq_zero_ext.semiring, + .. triv_sq_zero_ext.non_assoc_ring } + +instance [comm_monoid R] [add_comm_monoid M] + [distrib_mul_action R M] [distrib_mul_action Rᵐᵒᵖ M] [is_central_scalar R M] : + comm_monoid (tsze R M) := { mul_comm := λ x₁ x₂, ext (mul_comm x₁.1 x₂.1) $ - show x₁.1 • x₂.2 + x₂.1 • x₁.2 = x₂.1 • x₁.2 + x₁.1 • x₂.2, from add_comm _ _, + show x₁.1 • x₂.2 + op x₂.1 • x₁.2 = x₂.1 • x₁.2 + op x₁.1 • x₂.2, + by rw [op_smul_eq_smul, op_smul_eq_smul, add_comm] .. triv_sq_zero_ext.monoid } -instance [comm_semiring R] [add_comm_monoid M] [module R M] : comm_semiring (tsze R M) := +instance [comm_semiring R] [add_comm_monoid M] + [module R M] [module Rᵐᵒᵖ M] [is_central_scalar R M] : + comm_semiring (tsze R M) := { .. triv_sq_zero_ext.comm_monoid, .. triv_sq_zero_ext.non_assoc_semiring } -instance [comm_ring R] [add_comm_group M] [module R M] : comm_ring (tsze R M) := +instance [comm_ring R] [add_comm_group M] + [module R M] [module Rᵐᵒᵖ M] [is_central_scalar R M] : comm_ring (tsze R M) := { .. triv_sq_zero_ext.non_assoc_ring, .. triv_sq_zero_ext.comm_semiring } @@ -424,7 +515,7 @@ variables (R M) /-- The canonical inclusion of rings `R → triv_sq_zero_ext R M`. -/ @[simps apply] -def inl_hom [semiring R] [add_comm_monoid M] [module R M] : R →+* tsze R M := +def inl_hom [semiring R] [add_comm_monoid M] [module R M] [module Rᵐᵒᵖ M] : R →+* tsze R M := { to_fun := inl, map_one' := inl_one M, map_mul' := inl_mul M, @@ -434,75 +525,86 @@ def inl_hom [semiring R] [add_comm_monoid M] [module R M] : R →+* tsze R M := end mul section algebra -variables (S : Type*) (R : Type u) (M : Type v) -variables [comm_semiring S] [comm_semiring R] [add_comm_monoid M] -variables [algebra S R] [module S M] [module R M] [is_scalar_tower S R M] - -instance algebra' : algebra S (tsze R M) := +variables (S : Type*) (R R' : Type u) (M : Type v) +variables [comm_semiring S] [semiring R] [comm_semiring R'] [add_comm_monoid M] +variables [algebra S R] [algebra S R'] [module S M] +variables [module R M] [module Rᵐᵒᵖ M] [smul_comm_class R Rᵐᵒᵖ M] +variables [is_scalar_tower S R M] [is_scalar_tower S Rᵐᵒᵖ M] +variables [module R' M] [module R'ᵐᵒᵖ M] [is_central_scalar R' M] [is_scalar_tower S R' M] + +instance algebra' : algebra S (tsze R M) := { smul := (•), - commutes' := λ r x, mul_comm _ _, + commutes' := λ s x, ext (algebra.commutes _ _) $ + show algebra_map S R s • x.snd + op x.fst • 0 = x.fst • 0 + op (algebra_map S R s) • x.snd, + begin + rw [smul_zero, smul_zero, add_zero, zero_add], + rw [algebra.algebra_map_eq_smul_one, mul_opposite.op_smul, mul_opposite.op_one, + smul_assoc, one_smul, smul_assoc, one_smul], + end, smul_def' := λ r x, ext (algebra.smul_def _ _) $ - show r • x.2 = algebra_map S R r • x.2 + x.1 • 0, by rw [smul_zero, add_zero, algebra_map_smul], + show r • x.2 = algebra_map S R r • x.2 + op x.1 • 0, + by rw [smul_zero, add_zero, algebra_map_smul], .. (triv_sq_zero_ext.inl_hom R M).comp (algebra_map S R) } -- shortcut instance for the common case -instance : algebra R (tsze R M) := triv_sq_zero_ext.algebra' _ _ _ +instance : algebra R' (tsze R' M) := triv_sq_zero_ext.algebra' _ _ _ -lemma algebra_map_eq_inl : ⇑(algebra_map R (tsze R M)) = inl := rfl -lemma algebra_map_eq_inl_hom : algebra_map R (tsze R M) = inl_hom R M := rfl +lemma algebra_map_eq_inl : ⇑(algebra_map R' (tsze R' M)) = inl := rfl +lemma algebra_map_eq_inl_hom : algebra_map R' (tsze R' M) = inl_hom R' M := rfl lemma algebra_map_eq_inl' (s : S) : algebra_map S (tsze R M) s = inl (algebra_map S R s) := rfl /-- The canonical `R`-algebra projection `triv_sq_zero_ext R M → R`. -/ @[simps] -def fst_hom : tsze R M →ₐ[R] R := +def fst_hom : tsze R M →ₐ[S] R := { to_fun := fst, map_one' := fst_one, map_mul' := fst_mul, map_zero' := fst_zero, map_add' := fst_add, - commutes' := fst_inl M } + commutes' := λ r, fst_inl M _ } -variables {R S M} +variables {R R' S M} -lemma alg_hom_ext {A} [semiring A] [algebra R A] ⦃f g : tsze R M →ₐ[R] A⦄ +lemma alg_hom_ext {A} [semiring A] [algebra R' A] ⦃f g : tsze R' M →ₐ[R'] A⦄ (h : ∀ m, f (inr m) = g (inr m)) : f = g := alg_hom.to_linear_map_injective $ linear_map_ext (λ r, (f.commutes _).trans (g.commutes _).symm) h @[ext] -lemma alg_hom_ext' {A} [semiring A] [algebra R A] ⦃f g : tsze R M →ₐ[R] A⦄ - (h : f.to_linear_map.comp (inr_hom R M) = g.to_linear_map.comp (inr_hom R M)) : +lemma alg_hom_ext' {A} [semiring A] [algebra R' A] ⦃f g : tsze R' M →ₐ[R'] A⦄ + (h : f.to_linear_map.comp (inr_hom R' M) = g.to_linear_map.comp (inr_hom R' M)) : f = g := alg_hom_ext $ linear_map.congr_fun h -variables {A : Type*} [semiring A] [algebra R A] +variables {A : Type*} [semiring A] [algebra R' A] /-- There is an alg_hom from the trivial square zero extension to any `R`-algebra with a submodule whose products are all zero. See `triv_sq_zero_ext.lift` for this as an equiv. -/ -def lift_aux (f : M →ₗ[R] A) (hf : ∀ x y, f x * f y = 0) : tsze R M →ₐ[R] A := +def lift_aux (f : M →ₗ[R'] A) (hf : ∀ x y, f x * f y = 0) : tsze R' M →ₐ[R'] A := alg_hom.of_linear_map - ((algebra.linear_map _ _).comp (fst_hom R M).to_linear_map + f.comp (snd_hom R M)) - (show algebra_map R _ 1 + f (0 : M) = 1, by rw [map_zero, map_one, add_zero]) + ((algebra.linear_map _ _).comp (fst_hom R' R' M).to_linear_map + f.comp (snd_hom R' M)) + (show algebra_map R' _ 1 + f (0 : M) = 1, by rw [map_zero, map_one, add_zero]) (triv_sq_zero_ext.ind $ λ r₁ m₁, triv_sq_zero_ext.ind $ λ r₂ m₂, begin dsimp, - simp only [add_zero, zero_add, add_mul, mul_add, smul_mul_smul, hf, smul_zero], + simp only [add_zero, zero_add, add_mul, mul_add, smul_mul_smul, hf, smul_zero, + op_smul_eq_smul], rw [←ring_hom.map_mul, linear_map.map_add, ←algebra.commutes _ (f _), ←algebra.smul_def, ←algebra.smul_def, add_right_comm, add_assoc, linear_map.map_smul, linear_map.map_smul], end) -@[simp] lemma lift_aux_apply_inr (f : M →ₗ[R] A) (hf : ∀ x y, f x * f y = 0) (m : M) : +@[simp] lemma lift_aux_apply_inr (f : M →ₗ[R'] A) (hf : ∀ x y, f x * f y = 0) (m : M) : lift_aux f hf (inr m) = f m := -show algebra_map R A 0 + f m = f m, by rw [ring_hom.map_zero, zero_add] +show algebra_map R' A 0 + f m = f m, by rw [ring_hom.map_zero, zero_add] -@[simp] lemma lift_aux_comp_inr_hom (f : M →ₗ[R] A) (hf : ∀ x y, f x * f y = 0) : - (lift_aux f hf).to_linear_map.comp (inr_hom R M) = f := +@[simp] lemma lift_aux_comp_inr_hom (f : M →ₗ[R'] A) (hf : ∀ x y, f x * f y = 0) : + (lift_aux f hf).to_linear_map.comp (inr_hom R' M) = f := linear_map.ext $ lift_aux_apply_inr f hf /- When applied to `inr` itself, `lift_aux` is the identity. -/ @[simp] -lemma lift_aux_inr_hom : lift_aux (inr_hom R M) (inr_mul_inr R) = alg_hom.id R (tsze R M) := +lemma lift_aux_inr_hom : lift_aux (inr_hom R' M) (inr_mul_inr R') = alg_hom.id R' (tsze R' M) := alg_hom_ext' $ lift_aux_comp_inr_hom _ _ /-- A universal property of the trivial square-zero extension, providing a unique @@ -511,13 +613,16 @@ products. This isomorphism is named to match the very similar `complex.lift`. -/ @[simps] -def lift : {f : M →ₗ[R] A // ∀ x y, f x * f y = 0} ≃ (tsze R M →ₐ[R] A) := +def lift : {f : M →ₗ[R'] A // ∀ x y, f x * f y = 0} ≃ (tsze R' M →ₐ[R'] A) := { to_fun := λ f, lift_aux f f.prop, - inv_fun := λ F, ⟨F.to_linear_map.comp (inr_hom R M), λ x y, + inv_fun := λ F, ⟨F.to_linear_map.comp (inr_hom R' M), λ x y, (F.map_mul _ _).symm.trans $ (F.congr_arg $ inr_mul_inr _ _ _).trans F.map_zero⟩, left_inv := λ f, subtype.ext $ lift_aux_comp_inr_hom _ _, right_inv := λ F, alg_hom_ext' $ lift_aux_comp_inr_hom _ _, } +/-- This lemma is obviously simp-normal, but the linter times out while processing it.-/ +attribute [nolint simp_nf] lift_symm_apply_coe + end algebra end triv_sq_zero_ext diff --git a/src/analysis/normed_space/triv_sq_zero_ext.lean b/src/analysis/normed_space/triv_sq_zero_ext.lean index cf49d3b26b5a0..53f5ced3084a8 100644 --- a/src/analysis/normed_space/triv_sq_zero_ext.lean +++ b/src/analysis/normed_space/triv_sq_zero_ext.lean @@ -12,9 +12,6 @@ import topology.instances.triv_sq_zero_ext For now, this file contains results about `exp` for this type. -TODO: actually define a sensible norm on `triv_sq_zero_ext R M`, so that we have access to lemmas -like `exp_add`. - ## Main results * `triv_sq_zero_ext.fst_exp` @@ -22,6 +19,11 @@ like `exp_add`. * `triv_sq_zero_ext.exp_inl` * `triv_sq_zero_ext.exp_inr` +## TODO +* Actually define a sensible norm on `triv_sq_zero_ext R M`, so that we have access to lemmas + like `exp_add`. +* Generalize some of these results to non-commutative `R`. + -/ variables (𝕜 : Type*) {R M : Type*} @@ -36,8 +38,11 @@ variables [topological_space R] [topological_space M] /-- If `exp R x.fst` converges to `e` then `exp R x` converges to `inl e + inr (e • x.snd)`. -/ lemma has_sum_exp_series [field 𝕜] [char_zero 𝕜] [comm_ring R] - [add_comm_group M] [algebra 𝕜 R] [module R M] [module 𝕜 M] [is_scalar_tower 𝕜 R M] - [topological_ring R] [topological_add_group M] [has_continuous_smul R M] + [add_comm_group M] [algebra 𝕜 R] + [module R M] [module Rᵐᵒᵖ M] [is_central_scalar R M] + [module 𝕜 M] [is_scalar_tower 𝕜 R M] + [topological_ring R] [topological_add_group M] + [has_continuous_smul R M] (x : tsze R M) {e : R} (h : has_sum (λ n, exp_series 𝕜 R n (λ _, x.fst)) e) : has_sum (λ n, exp_series 𝕜 (tsze R M) n (λ _, x)) (inl e + inr (e • x.snd)) := begin @@ -62,7 +67,8 @@ end topology section normed_ring variables [is_R_or_C 𝕜] [normed_comm_ring R] [add_comm_group M] -variables [normed_algebra 𝕜 R] [module R M] [module 𝕜 M] [is_scalar_tower 𝕜 R M] +variables [normed_algebra 𝕜 R] [module R M] [module Rᵐᵒᵖ M] [is_central_scalar R M] +variables [module 𝕜 M] [is_scalar_tower 𝕜 R M] variables [topological_space M] [topological_ring R] variables [topological_add_group M] [has_continuous_smul R M] variables [complete_space R] [t2_space R] [t2_space M] @@ -97,7 +103,8 @@ end normed_ring section normed_field variables [is_R_or_C 𝕜] [normed_field R] [add_comm_group M] -variables [normed_algebra 𝕜 R] [module R M] [module 𝕜 M] [is_scalar_tower 𝕜 R M] +variables [normed_algebra 𝕜 R] [module R M] [module Rᵐᵒᵖ M] [is_central_scalar R M] +variables [module 𝕜 M] [is_scalar_tower 𝕜 R M] variables [topological_space M] [topological_ring R] variables [topological_add_group M] [has_continuous_smul R M] variables [complete_space R] [t2_space R] [t2_space M] diff --git a/src/linear_algebra/exterior_algebra/basic.lean b/src/linear_algebra/exterior_algebra/basic.lean index fba425a184892..a5bc276ee5af1 100644 --- a/src/linear_algebra/exterior_algebra/basic.lean +++ b/src/linear_algebra/exterior_algebra/basic.lean @@ -160,10 +160,11 @@ variables {M} /-- The canonical map from `exterior_algebra R M` into `triv_sq_zero_ext R M` that sends `exterior_algebra.ι` to `triv_sq_zero_ext.inr`. -/ -def to_triv_sq_zero_ext : exterior_algebra R M →ₐ[R] triv_sq_zero_ext R M := +def to_triv_sq_zero_ext [module Rᵐᵒᵖ M] [is_central_scalar R M] : + exterior_algebra R M →ₐ[R] triv_sq_zero_ext R M := lift R ⟨triv_sq_zero_ext.inr_hom R M, λ m, triv_sq_zero_ext.inr_mul_inr R m m⟩ -@[simp] lemma to_triv_sq_zero_ext_ι (x : M) : +@[simp] lemma to_triv_sq_zero_ext_ι [module Rᵐᵒᵖ M] [is_central_scalar R M] (x : M) : to_triv_sq_zero_ext (ι R x) = triv_sq_zero_ext.inr x := lift_ι_apply _ _ _ _ @@ -172,7 +173,11 @@ lift_ι_apply _ _ _ _ As an implementation detail, we implement this using `triv_sq_zero_ext` which has a suitable algebra structure. -/ def ι_inv : exterior_algebra R M →ₗ[R] M := -(triv_sq_zero_ext.snd_hom R M).comp to_triv_sq_zero_ext.to_linear_map +begin + letI : module Rᵐᵒᵖ M := module.comp_hom _ ((ring_hom.id R).from_opposite mul_comm), + haveI : is_central_scalar R M := ⟨λ r m, rfl⟩, + exact (triv_sq_zero_ext.snd_hom R M).comp to_triv_sq_zero_ext.to_linear_map +end lemma ι_left_inverse : function.left_inverse ι_inv (ι R : M → exterior_algebra R M) := λ x, by simp [ι_inv] @@ -190,7 +195,9 @@ by rw [←ι_inj R x 0, linear_map.map_zero] @[simp] lemma ι_eq_algebra_map_iff (x : M) (r : R) : ι R x = algebra_map R _ r ↔ x = 0 ∧ r = 0 := begin refine ⟨λ h, _, _⟩, - { have hf0 : to_triv_sq_zero_ext (ι R x) = (0, x), from to_triv_sq_zero_ext_ι _, + { letI : module Rᵐᵒᵖ M := module.comp_hom _ ((ring_hom.id R).from_opposite mul_comm), + haveI : is_central_scalar R M := ⟨λ r m, rfl⟩, + have hf0 : to_triv_sq_zero_ext (ι R x) = (0, x), from to_triv_sq_zero_ext_ι _, rw [h, alg_hom.commutes] at hf0, have : r = 0 ∧ 0 = x := prod.ext_iff.1 hf0, exact this.symm.imp_left eq.symm, }, diff --git a/src/linear_algebra/tensor_algebra/basic.lean b/src/linear_algebra/tensor_algebra/basic.lean index 2fe7d35868349..98229da93f802 100644 --- a/src/linear_algebra/tensor_algebra/basic.lean +++ b/src/linear_algebra/tensor_algebra/basic.lean @@ -188,10 +188,11 @@ variables {M} /-- The canonical map from `tensor_algebra R M` into `triv_sq_zero_ext R M` that sends `tensor_algebra.ι` to `triv_sq_zero_ext.inr`. -/ -def to_triv_sq_zero_ext : tensor_algebra R M →ₐ[R] triv_sq_zero_ext R M := +def to_triv_sq_zero_ext [module Rᵐᵒᵖ M] [is_central_scalar R M] : + tensor_algebra R M →ₐ[R] triv_sq_zero_ext R M := lift R (triv_sq_zero_ext.inr_hom R M) -@[simp] lemma to_triv_sq_zero_ext_ι (x : M) : +@[simp] lemma to_triv_sq_zero_ext_ι (x : M) [module Rᵐᵒᵖ M] [is_central_scalar R M] : to_triv_sq_zero_ext (ι R x) = triv_sq_zero_ext.inr x := lift_ι_apply _ _ @@ -200,7 +201,11 @@ lift_ι_apply _ _ As an implementation detail, we implement this using `triv_sq_zero_ext` which has a suitable algebra structure. -/ def ι_inv : tensor_algebra R M →ₗ[R] M := -(triv_sq_zero_ext.snd_hom R M).comp to_triv_sq_zero_ext.to_linear_map +begin + letI : module Rᵐᵒᵖ M := module.comp_hom _ ((ring_hom.id R).from_opposite mul_comm), + haveI : is_central_scalar R M := ⟨λ r m, rfl⟩, + exact (triv_sq_zero_ext.snd_hom R M).comp to_triv_sq_zero_ext.to_linear_map +end lemma ι_left_inverse : function.left_inverse ι_inv (ι R : M → tensor_algebra R M) := λ x, by simp [ι_inv] @@ -218,7 +223,9 @@ variables {R} @[simp] lemma ι_eq_algebra_map_iff (x : M) (r : R) : ι R x = algebra_map R _ r ↔ x = 0 ∧ r = 0 := begin refine ⟨λ h, _, _⟩, - { have hf0 : to_triv_sq_zero_ext (ι R x) = (0, x), from lift_ι_apply _ _, + { letI : module Rᵐᵒᵖ M := module.comp_hom _ ((ring_hom.id R).from_opposite mul_comm), + haveI : is_central_scalar R M := ⟨λ r m, rfl⟩, + have hf0 : to_triv_sq_zero_ext (ι R x) = (0, x), from lift_ι_apply _ _, rw [h, alg_hom.commutes] at hf0, have : r = 0 ∧ 0 = x := prod.ext_iff.1 hf0, exact this.symm.imp_left eq.symm, }, diff --git a/src/topology/instances/triv_sq_zero_ext.lean b/src/topology/instances/triv_sq_zero_ext.lean index 9f4c07efe72c8..c791c70b0f1dd 100644 --- a/src/topology/instances/triv_sq_zero_ext.lean +++ b/src/topology/instances/triv_sq_zero_ext.lean @@ -89,14 +89,15 @@ instance [has_add R] [has_add M] has_continuous_add (tsze R M) := prod.has_continuous_add -instance [has_mul R] [has_add M] [has_smul R M] - [has_continuous_mul R] [has_continuous_smul R M] [has_continuous_add M] : +instance [has_mul R] [has_add M] [has_smul R M] [has_smul Rᵐᵒᵖ M] + [has_continuous_mul R] [has_continuous_smul R M] [has_continuous_smul Rᵐᵒᵖ M] + [has_continuous_add M] : has_continuous_mul (tsze R M) := ⟨((continuous_fst.comp _root_.continuous_fst).mul (continuous_fst.comp _root_.continuous_snd)) .prod_mk $ ((continuous_fst.comp _root_.continuous_fst).smul (continuous_snd.comp _root_.continuous_snd)).add - ((continuous_fst.comp _root_.continuous_snd).smul + ((mul_opposite.continuous_op.comp $ continuous_fst.comp $ _root_.continuous_snd).smul (continuous_snd.comp _root_.continuous_fst))⟩ instance [has_neg R] [has_neg M] @@ -104,13 +105,18 @@ instance [has_neg R] [has_neg M] has_continuous_neg (tsze R M) := prod.has_continuous_neg -instance [semiring R] [add_comm_monoid M] [module R M] - [topological_semiring R] [has_continuous_add M] [has_continuous_smul R M] : - topological_semiring (tsze R M) := +/-- This is not an instance due to complaints by the `fails_quickly` linter. At any rate, we only +really care about the `topological_ring` instance below. -/ +lemma topological_semiring [semiring R] [add_comm_monoid M] [module R M] [module Rᵐᵒᵖ M] + [topological_semiring R] [has_continuous_add M] + [has_continuous_smul R M] [has_continuous_smul Rᵐᵒᵖ M] : + -- note: lean times out looking for the non_assoc_semiring instance without this hint + @topological_semiring (tsze R M) _ (non_assoc_semiring.to_non_unital_non_assoc_semiring _) := {} -instance [comm_ring R] [add_comm_group M] [module R M] - [topological_ring R] [topological_add_group M] [has_continuous_smul R M] : +instance [ring R] [add_comm_group M] [module R M] [module Rᵐᵒᵖ M] + [topological_ring R] [topological_add_group M] + [has_continuous_smul R M] [has_continuous_smul Rᵐᵒᵖ M] : topological_ring (tsze R M) := {} From 22131150f88a2d125713ffa0f4693e3355b1eb49 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 23 Feb 2023 11:46:51 +0000 Subject: [PATCH 0515/1291] feat(topology/uniform_space/cauchy): `mul_opposite X` is complete when `X` is (#18471) The new import is needed because the lemmas about the topology on `mul_opposite` aren't available yet, even though the instance itself is (via the uniformity instance). Forward-ported as https://github.com/leanprover-community/mathlib4/pull/2404. --- src/topology/uniform_space/cauchy.lean | 7 +++++++ 1 file changed, 7 insertions(+) diff --git a/src/topology/uniform_space/cauchy.lean b/src/topology/uniform_space/cauchy.lean index b96f8ef30b9b6..f3eb455fbe77b 100644 --- a/src/topology/uniform_space/cauchy.lean +++ b/src/topology/uniform_space/cauchy.lean @@ -3,6 +3,7 @@ Copyright (c) 2017 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro -/ +import topology.algebra.constructions import topology.bases import topology.uniform_space.basic /-! @@ -371,6 +372,12 @@ instance complete_space.prod [uniform_space β] [complete_space α] [complete_sp from filter.le_lift.2 (λ s hs, filter.le_lift'.2 $ λ t ht, inter_mem (hx1 hs) (hx2 ht))⟩ } +@[to_additive] +instance complete_space.mul_opposite [complete_space α] : complete_space αᵐᵒᵖ := +{ complete := λ f hf, mul_opposite.op_surjective.exists.mpr $ + let ⟨x, hx⟩ := complete_space.complete (hf.map mul_opposite.uniform_continuous_unop) in + ⟨x, (map_le_iff_le_comap.mp hx).trans_eq $ mul_opposite.comap_unop_nhds _⟩} + /--If `univ` is complete, the space is a complete space -/ lemma complete_space_of_is_complete_univ (h : is_complete (univ : set α)) : complete_space α := ⟨λ f hf, let ⟨x, _, hx⟩ := h f hf ((@principal_univ α).symm ▸ le_top) in ⟨x, hx⟩⟩ From 651f93df29be483df3eb3778e1500fbf18bac41e Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Thu, 23 Feb 2023 22:53:04 +0000 Subject: [PATCH 0516/1291] feat(analysis/normed_space/star/multiplier): construct the multiplier algebra of a C*-algebra (#15869) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Define the multiplier algebra of a C⋆-algebra as the algebra (over `𝕜`) of double centralizers, for which we provide the localized notation `𝓜(𝕜, A)`. A double centralizer is a pair of continuous linear maps `L R : A →L[𝕜] A` satisfying the intertwining condition `R x * y = x * L y`. There is a natural embedding `A → 𝓜(𝕜, A)` which sends `a : A` to the continuous linear maps `L R : A →L[𝕜] A` given by left and right multiplication by `a`, and we provide this map as a coercion. In this PR we put all the natural structures on `𝓜(𝕜, A)`, culminating in the fact that it is a unital C⋆-algebra when A is a (unital or non-unital) C⋆-algebra. Co-authored-by: Jon Bannon Co-authored-by: Jon Bannon --- .../normed_space/star/multiplier.lean | 512 ++++++++++++++++++ 1 file changed, 512 insertions(+) create mode 100644 src/analysis/normed_space/star/multiplier.lean diff --git a/src/analysis/normed_space/star/multiplier.lean b/src/analysis/normed_space/star/multiplier.lean new file mode 100644 index 0000000000000..a72a923aeb97a --- /dev/null +++ b/src/analysis/normed_space/star/multiplier.lean @@ -0,0 +1,512 @@ +/- +Copyright (c) 2022 Jireh Loreaux. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jireh Loreaux, Jon Bannon +-/ + +import algebra.star.star_alg_hom +import analysis.normed_space.star.basic +import analysis.normed_space.operator_norm +import analysis.special_functions.pow +import analysis.normed_space.star.mul + +/-! +# Multiplier Algebra of a C⋆-algebra + +Define the multiplier algebra of a C⋆-algebra as the algebra (over `𝕜`) of double centralizers, +for which we provide the localized notation `𝓜(𝕜, A)`. A double centralizer is a pair of +continuous linear maps `L R : A →L[𝕜] A` satisfying the intertwining condition `R x * y = x * L y`. + +There is a natural embedding `A → 𝓜(𝕜, A)` which sends `a : A` to the continuous linear maps +`L R : A →L[𝕜] A` given by left and right multiplication by `a`, and we provide this map as a +coercion. + +The multiplier algebra corresponds to a non-commutative Stone–Čech compactification in the sense +that when the algebra `A` is commutative, it can be identified with `C₀(X, ℂ)` for some locally +compact Hausdorff space `X`, and in that case `𝓜(𝕜, A)` can be identified with `C(β X, ℂ)`. + +## Implementation notes + +We make the hypotheses on `𝕜` as weak as possible so that, in particular, this construction works +for both `𝕜 = ℝ` and `𝕜 = ℂ`. + +The reader familiar with C⋆-algebra theory may recognize that one +only needs `L` and `R` to be functions instead of continuous linear maps, at least when `A` is a +C⋆-algebra. Our intention is simply to eventually provide a constructor for this situation. + +We pull back the `normed_algebra` structure (and everything contained therein) through the +ring (even algebra) homomorphism +`double_centralizer.to_prod_mul_opposite_hom : 𝓜(𝕜, A) →+* (A →L[𝕜] A) × (A →L[𝕜] A)ᵐᵒᵖ` which +sends `a : 𝓜(𝕜, A)` to `(a.fst, mul_opposite.op a.snd)`. The star structure is provided +separately. + +## References + +* https://en.wikipedia.org/wiki/Multiplier_algebra + +## TODO + ++ Define a type synonym for `𝓜(𝕜, A)` which is equipped with the strict uniform space structure + and show it is complete ++ Show that the image of `A` in `𝓜(𝕜, A)` is an essential ideal ++ Prove the universal property of `𝓜(𝕜, A)` ++ Construct a double centralizer from a pair of maps (not necessarily linear or continuous) + `L : A → A`, `R : A → A` satisfying the centrality condition `∀ x y, R x * y = x * L y`. ++ Show that if `A` is unital, then `A ≃⋆ₐ[𝕜] 𝓜(𝕜, A)`. +-/ + +open_locale nnreal ennreal +open nnreal continuous_linear_map mul_opposite + +universes u v + +/-- The type of *double centralizers*, also known as the *multiplier algebra* and denoted by +`𝓜(𝕜, A)`, of a non-unital normed algebra. + +If `x : 𝓜(𝕜, A)`, then `x.fst` and `x.snd` are what is usually referred to as $L$ and $R$. -/ +@[ext] +structure double_centralizer (𝕜 : Type u) (A : Type v) [nontrivially_normed_field 𝕜] + [non_unital_normed_ring A] [normed_space 𝕜 A] [smul_comm_class 𝕜 A A] [is_scalar_tower 𝕜 A A] + extends (A →L[𝕜] A) × (A →L[𝕜] A) := +(central : ∀ x y : A, snd x * y = x * fst y) + +localized "notation `𝓜(` 𝕜 `, ` A `)` := double_centralizer 𝕜 A" in multiplier_algebra + +namespace double_centralizer + +section nontrivially_normed + +variables (𝕜 A : Type*) [nontrivially_normed_field 𝕜] [non_unital_normed_ring A] +variables [normed_space 𝕜 A] [smul_comm_class 𝕜 A A] [is_scalar_tower 𝕜 A A] + +/-! +### Algebraic structure + +Because the multiplier algebra is defined as the algebra of double centralizers, there is a natural +injection `double_centralizer.to_prod_mul_opposite : 𝓜(𝕜, A) → (A →L[𝕜] A) × (A →L[𝕜] A)ᵐᵒᵖ` +defined by `λ a, (a.fst, mul_opposite.op a.snd)`. We use this map to pull back the ring, module and +algebra structure from `(A →L[𝕜] A) × (A →L[𝕜] A)ᵐᵒᵖ` to `𝓜(𝕜, A)`. -/ + +variables {𝕜 A} + +lemma range_to_prod : set.range to_prod = {lr : (A →L[𝕜] A) × _ | ∀ x y, lr.2 x * y = x * lr.1 y} := +set.ext $ λ x, ⟨by {rintro ⟨a, rfl⟩, exact a.central}, λ hx, ⟨⟨x, hx⟩, rfl⟩⟩ + +instance : has_add 𝓜(𝕜, A) := +{ add := λ a b, + { to_prod := a.to_prod + b.to_prod, + central := λ x y, show (a.snd + b.snd) x * y = x * (a.fst + b.fst) y, + by simp only [continuous_linear_map.add_apply, mul_add, add_mul, central] } } + +instance : has_zero 𝓜(𝕜, A) := +{ zero := + { to_prod := 0, + central := λ x y, (zero_mul y).trans (mul_zero x).symm } } + +instance : has_neg 𝓜(𝕜, A) := +{ neg := λ a, + { to_prod := -a.to_prod, + central := λ x y, show -a.snd x * y = x * -a.fst y, + by simp only [continuous_linear_map.neg_apply, neg_mul, mul_neg, central] } } + +instance : has_sub 𝓜(𝕜, A) := +{ sub := λ a b, + { to_prod := a.to_prod - b.to_prod, + central := λ x y, show (a.snd - b.snd) x * y = x * (a.fst - b.fst) y, + by simp only [continuous_linear_map.sub_apply, sub_mul, mul_sub, central] } } + +section scalars + +variables {S : Type*} [monoid S] [distrib_mul_action S A] [smul_comm_class 𝕜 S A] + [has_continuous_const_smul S A] [is_scalar_tower S A A] [smul_comm_class S A A] + +instance : has_smul S 𝓜(𝕜, A) := +{ smul := λ s a, + { to_prod := s • a.to_prod, + central := λ x y, show (s • a.snd) x * y = x * (s • a.fst) y, + by simp only [continuous_linear_map.smul_apply, mul_smul_comm, smul_mul_assoc, central] } } + +@[simp] lemma smul_to_prod (s : S) (a : 𝓜(𝕜, A)) : (s • a).to_prod = s • a.to_prod := rfl +lemma smul_fst (s : S) (a : 𝓜(𝕜, A)) : (s • a).fst = s • a.fst := rfl +lemma smul_snd (s : S) (a : 𝓜(𝕜, A)) : (s • a).snd = s • a.snd := rfl + +variables {T : Type*} [monoid T] [distrib_mul_action T A] [smul_comm_class 𝕜 T A] + [has_continuous_const_smul T A] [is_scalar_tower T A A] [smul_comm_class T A A] + +instance [has_smul S T] [is_scalar_tower S T A] : is_scalar_tower S T 𝓜(𝕜, A) := +{ smul_assoc := λ _ _ a, ext _ _ $ smul_assoc _ _ a.to_prod } + +instance [smul_comm_class S T A] : smul_comm_class S T 𝓜(𝕜, A) := +{ smul_comm := λ _ _ a, ext _ _ $ smul_comm _ _ a.to_prod } + +instance {R : Type*} [semiring R] [module R A] [smul_comm_class 𝕜 R A] + [has_continuous_const_smul R A] [is_scalar_tower R A A] [smul_comm_class R A A] + [module Rᵐᵒᵖ A] [is_central_scalar R A] : is_central_scalar R 𝓜(𝕜, A) := +{ op_smul_eq_smul := λ _ a, ext _ _ $ op_smul_eq_smul _ a.to_prod } + +end scalars + +instance : has_one 𝓜(𝕜, A) := ⟨⟨1, λ x y, rfl⟩⟩ + +instance : has_mul 𝓜(𝕜, A) := +{ mul := λ a b, + { to_prod := (a.fst.comp b.fst, b.snd.comp a.snd), + central := λ x y, show b.snd (a.snd x) * y = x * a.fst (b.fst y), + by simp only [central] } } + +instance : has_nat_cast 𝓜(𝕜, A) := +{ nat_cast := λ n, ⟨n, λ x y, + begin + rw [prod.snd_nat_cast, prod.fst_nat_cast], + simp only [←nat.smul_one_eq_coe, smul_apply, one_apply, mul_smul_comm, smul_mul_assoc], + end⟩ } + +instance : has_int_cast 𝓜(𝕜, A) := +{ int_cast := λ n, ⟨n, λ x y, + begin + rw [prod.snd_int_cast, prod.fst_int_cast], + simp only [←int.smul_one_eq_coe, smul_apply, one_apply, mul_smul_comm, smul_mul_assoc], + end⟩ } + +instance : has_pow 𝓜(𝕜, A) ℕ := +{ pow := λ a n, ⟨a.to_prod ^ n, λ x y, + begin + induction n with k hk generalizing x y, + { refl }, + { rw [prod.pow_snd, prod.pow_fst] at hk ⊢, + rw [pow_succ a.snd, mul_apply, a.central, hk, pow_succ' a.fst, mul_apply] }, + end⟩ } + +instance : inhabited 𝓜(𝕜, A) := ⟨0⟩ + +@[simp] lemma add_to_prod (a b : 𝓜(𝕜, A)) : (a + b).to_prod = a.to_prod + b.to_prod := rfl +@[simp] lemma zero_to_prod : (0 : 𝓜(𝕜, A)).to_prod = 0 := rfl +@[simp] lemma neg_to_prod (a : 𝓜(𝕜, A)) : (-a).to_prod = -a.to_prod := rfl +@[simp] lemma sub_to_prod (a b : 𝓜(𝕜, A)) : (a - b).to_prod = a.to_prod - b.to_prod := rfl +@[simp] lemma one_to_prod : (1 : 𝓜(𝕜, A)).to_prod = 1 := rfl +@[simp] lemma nat_cast_to_prod (n : ℕ) : (n : 𝓜(𝕜 , A)).to_prod = n := rfl +@[simp] lemma int_cast_to_prod (n : ℤ) : (n : 𝓜(𝕜 , A)).to_prod = n := rfl +@[simp] lemma pow_to_prod (n : ℕ) (a : 𝓜(𝕜, A)) : (a ^ n).to_prod = a.to_prod ^ n := rfl + +lemma add_fst (a b : 𝓜(𝕜, A)) : (a + b).fst = a.fst + b.fst := rfl +lemma add_snd (a b : 𝓜(𝕜, A)) : (a + b).snd = a.snd + b.snd := rfl +lemma zero_fst : (0 : 𝓜(𝕜, A)).fst = 0 := rfl +lemma zero_snd : (0 : 𝓜(𝕜, A)).snd = 0 := rfl +lemma neg_fst (a : 𝓜(𝕜, A)) : (-a).fst = -a.fst := rfl +lemma neg_snd (a : 𝓜(𝕜, A)) : (-a).snd = -a.snd := rfl +lemma sub_fst (a b : 𝓜(𝕜, A)) : (a - b).fst = a.fst - b.fst := rfl +lemma sub_snd (a b : 𝓜(𝕜, A)) : (a - b).snd = a.snd - b.snd := rfl +lemma one_fst : (1 : 𝓜(𝕜, A)).fst = 1 := rfl +lemma one_snd : (1 : 𝓜(𝕜, A)).snd = 1 := rfl +@[simp] lemma mul_fst (a b : 𝓜(𝕜, A)) : (a * b).fst = a.fst * b.fst := rfl +@[simp] lemma mul_snd (a b : 𝓜(𝕜, A)) : (a * b).snd = b.snd * a.snd := rfl +lemma nat_cast_fst (n : ℕ) : (n : 𝓜(𝕜 , A)).fst = n := rfl +lemma nat_cast_snd (n : ℕ) : (n : 𝓜(𝕜 , A)).snd = n := rfl +lemma int_cast_fst (n : ℤ) : (n : 𝓜(𝕜 , A)).fst = n := rfl +lemma int_cast_snd (n : ℤ) : (n : 𝓜(𝕜 , A)).snd = n := rfl +lemma pow_fst (n : ℕ) (a : 𝓜(𝕜, A)) : (a ^ n).fst = a.fst ^ n := rfl +lemma pow_snd (n : ℕ) (a : 𝓜(𝕜, A)) : (a ^ n).snd = a.snd ^ n := rfl + +/-- The natural injection from `double_centralizer.to_prod` except the second coordinate inherits +`mul_opposite.op`. The ring structure on `𝓜(𝕜, A)` is the pullback under this map. -/ +def to_prod_mul_opposite : 𝓜(𝕜, A) → (A →L[𝕜] A) × (A →L[𝕜] A)ᵐᵒᵖ := +λ a, (a.fst, mul_opposite.op a.snd) + +lemma to_prod_mul_opposite_injective : + function.injective (to_prod_mul_opposite : 𝓜(𝕜, A) → (A →L[𝕜] A) × (A →L[𝕜] A)ᵐᵒᵖ) := +λ a b h, let h' := prod.ext_iff.mp h in ext _ _ $ prod.ext h'.1 $ mul_opposite.op_injective h'.2 + +lemma range_to_prod_mul_opposite : + set.range to_prod_mul_opposite = {lr : (A →L[𝕜] A) × _ | ∀ x y, unop lr.2 x * y = x * lr.1 y} := +set.ext $ λ x, + ⟨by {rintro ⟨a, rfl⟩, exact a.central}, λ hx, ⟨⟨(x.1, unop x.2), hx⟩, prod.ext rfl rfl⟩⟩ + +/-- The ring structure is inherited as the pullback under the injective map +`double_centralizer.to_prod_mul_opposite : 𝓜(𝕜, A) → (A →L[𝕜] A) × (A →L[𝕜] A)ᵐᵒᵖ` -/ +instance : ring 𝓜(𝕜, A) := +to_prod_mul_opposite_injective.ring _ + rfl rfl (λ _ _, rfl) (λ _ _, rfl) (λ _, rfl) (λ _ _, rfl) + (λ x n, prod.ext rfl $ mul_opposite.op_smul _ _) + (λ x n, prod.ext rfl $ mul_opposite.op_smul _ _) + (λ x n, prod.ext rfl $ mul_opposite.op_pow _ _) + (λ _, rfl) (λ _, rfl) + +/-- The canonical map `double_centralizer.to_prod` as an additive group homomorphism. -/ +@[simps] +def to_prod_hom : 𝓜(𝕜, A) →+ (A →L[𝕜] A) × (A →L[𝕜] A) := +{ to_fun := to_prod, + map_zero' := rfl, + map_add' := λ x y, rfl } + +/-- The canonical map `double_centralizer.to_prod_mul_opposite` as a ring homomorphism. -/ +@[simps] +def to_prod_mul_opposite_hom : 𝓜(𝕜, A) →+* (A →L[𝕜] A) × (A →L[𝕜] A)ᵐᵒᵖ := +{ to_fun := to_prod_mul_opposite, + map_zero' := rfl, + map_one' := rfl, + map_add' := λ x y, rfl, + map_mul' := λ x y, rfl } + +/-- The module structure is inherited as the pullback under the additive group monomorphism +`double_centralizer.to_prod : 𝓜(𝕜, A) →+ (A →L[𝕜] A) × (A →L[𝕜] A)` -/ +instance {S : Type*} [semiring S] [module S A] [smul_comm_class 𝕜 S A] + [has_continuous_const_smul S A] [is_scalar_tower S A A] [smul_comm_class S A A] : + module S 𝓜(𝕜, A) := +function.injective.module S to_prod_hom ext (λ x y, rfl) + +-- TODO: generalize to `algebra S 𝓜(𝕜, A)` once `continuous_linear_map.algebra` is generalized. +instance : algebra 𝕜 𝓜(𝕜, A) := +{ to_fun := λ k, + { to_prod := algebra_map 𝕜 ((A →L[𝕜] A) × (A →L[𝕜] A)) k, + central := λ x y, by simp_rw [prod.algebra_map_apply, algebra.algebra_map_eq_smul_one, + smul_apply, one_apply, mul_smul_comm, smul_mul_assoc] }, + map_one' := ext _ _ $ map_one $ algebra_map 𝕜 ((A →L[𝕜] A) × (A →L[𝕜] A)), + map_mul' := λ k₁ k₂, ext _ _ $ prod.ext (map_mul (algebra_map 𝕜 (A →L[𝕜] A)) _ _) + ((map_mul (algebra_map 𝕜 (A →L[𝕜] A)) _ _).trans (algebra.commutes _ _)), + map_zero' := ext _ _ $ map_zero $ algebra_map 𝕜 ((A →L[𝕜] A) × (A →L[𝕜] A)), + map_add' := λ _ _, ext _ _ $ map_add (algebra_map 𝕜 ((A →L[𝕜] A) × (A →L[𝕜] A))) _ _, + commutes' := λ _ _, ext _ _ $ prod.ext (algebra.commutes _ _) (algebra.commutes _ _).symm, + smul_def' := λ _ _, ext _ _ $ prod.ext (algebra.smul_def _ _) + ((algebra.smul_def _ _).trans $ algebra.commutes _ _) } + +@[simp] lemma algebra_map_to_prod (k : 𝕜) : + (algebra_map 𝕜 𝓜(𝕜, A) k).to_prod = algebra_map 𝕜 _ k := rfl +lemma algebra_map_fst (k : 𝕜) : (algebra_map 𝕜 𝓜(𝕜, A) k).fst = algebra_map 𝕜 _ k := rfl +lemma algebra_map_snd (k : 𝕜) : (algebra_map 𝕜 𝓜(𝕜, A) k).snd = algebra_map 𝕜 _ k := rfl + +/-! +### Star structure +-/ + +section star + +variables [star_ring 𝕜] [star_ring A] [star_module 𝕜 A] [normed_star_group A] + +/-- The star operation on `a : 𝓜(𝕜, A)` is given by +`(star a).to_prod = (star ∘ a.snd ∘ star, star ∘ a.fst ∘ star)`. -/ +instance : has_star 𝓜(𝕜, A) := +{ star := λ a, + { fst := (((starₗᵢ 𝕜 : A ≃ₗᵢ⋆[𝕜] A) : A →L⋆[𝕜] A).comp a.snd).comp + ((starₗᵢ 𝕜 : A ≃ₗᵢ⋆[𝕜] A) : A →L⋆[𝕜] A), + snd := (((starₗᵢ 𝕜 : A ≃ₗᵢ⋆[𝕜] A) : A →L⋆[𝕜] A).comp a.fst).comp + ((starₗᵢ 𝕜 : A ≃ₗᵢ⋆[𝕜] A) : A →L⋆[𝕜] A), + central := λ x y, by simpa only [star_mul, star_star] + using (congr_arg star (a.central (star y) (star x))).symm } } + +@[simp] lemma star_fst (a : 𝓜(𝕜, A)) (b : A) : (star a).fst b = star (a.snd (star b)) := rfl +@[simp] lemma star_snd (a : 𝓜(𝕜, A)) (b : A) : (star a).snd b = star (a.fst (star b)) := rfl + +instance : star_add_monoid 𝓜(𝕜, A) := +{ star_involutive := λ x, by {ext; simp only [star_fst, star_snd, star_star]}, + star_add := λ x y, by {ext; simp only [star_fst, star_snd, add_fst, add_snd, + continuous_linear_map.add_apply, star_add]}, + .. double_centralizer.has_star } + +instance : star_ring 𝓜(𝕜, A) := +{ star_mul := λ a b, by {ext; simp only [star_fst, star_snd, mul_fst, mul_snd, star_star, + continuous_linear_map.coe_mul, function.comp_app]}, + .. double_centralizer.star_add_monoid } + +instance : star_module 𝕜 𝓜(𝕜, A) := +{ star_smul := λ k a, by {ext; exact star_smul _ _}, + .. double_centralizer.star_add_monoid } + +end star + +/-! +### Coercion from an algebra into its multiplier algebra +-/ + +/-- The natural coercion of `A` into `𝓜(𝕜, A)` given by sending `a : A` to the pair of linear +maps `Lₐ Rₐ : A →L[𝕜] A` given by left- and right-multiplication by `a`, respectively. + +Warning: if `A = 𝕜`, then this is a coercion which is not definitionally equal to the +`algebra_map 𝕜 𝓜(𝕜, 𝕜)` coercion, but these are propositionally equal. See +`double_centralizer.coe_eq_algebra_map` below. -/ +noncomputable instance : has_coe_t A 𝓜(𝕜, A) := +{ coe := λ a, + { fst := continuous_linear_map.mul 𝕜 A a, + snd := (continuous_linear_map.mul 𝕜 A).flip a, + central := λ x y, mul_assoc _ _ _ } } + +@[simp, norm_cast] +lemma coe_fst (a : A) : (a : 𝓜(𝕜, A)).fst = continuous_linear_map.mul 𝕜 A a := rfl +@[simp, norm_cast] +lemma coe_snd (a : A) : (a : 𝓜(𝕜, A)).snd = (continuous_linear_map.mul 𝕜 A).flip a := rfl + +lemma coe_eq_algebra_map : (coe : 𝕜 → 𝓜(𝕜, 𝕜)) = algebra_map 𝕜 𝓜(𝕜, 𝕜) := +begin + ext; + simp only [coe_fst, mul_apply', mul_one, algebra_map_to_prod, prod.algebra_map_apply, coe_snd, + flip_apply, one_mul]; + simp only [algebra.algebra_map_eq_smul_one, smul_apply, one_apply, smul_eq_mul, mul_one], +end + +/-- The coercion of an algebra into its multiplier algebra as a non-unital star algebra +homomorphism. -/ +@[simps] +noncomputable def coe_hom [star_ring 𝕜] [star_ring A] [star_module 𝕜 A] [normed_star_group A] : + A →⋆ₙₐ[𝕜] 𝓜(𝕜, A) := +{ to_fun := λ a, a, + map_smul' := λ k a, by ext; simp only [coe_fst, coe_snd, continuous_linear_map.map_smul, + smul_fst, smul_snd], + map_zero' := by ext; simp only [coe_fst, coe_snd, map_zero, zero_fst, zero_snd], + map_add' := λ a b, by ext; simp only [coe_fst, coe_snd, map_add, add_fst, add_snd], + map_mul' := λ a b, by ext; simp only [coe_fst, coe_snd, mul_apply', flip_apply, mul_fst, mul_snd, + continuous_linear_map.coe_mul, function.comp_app, mul_assoc], + map_star' := λ a, by ext; simp only [coe_fst, coe_snd, mul_apply', star_fst, star_snd, + flip_apply, star_mul, star_star] } + +/-! +### Norm structures +We define the norm structure on `𝓜(𝕜, A)` as the pullback under +`double_centralizer.to_prod_mul_opposite_hom : 𝓜(𝕜, A) →+* (A →L[𝕜] A) × (A →L[𝕜] A)ᵐᵒᵖ`, which +provides a definitional isometric embedding. Consequently, completeness of `𝓜(𝕜, A)` is obtained +by proving that the range of this map is closed. + +In addition, we prove that `𝓜(𝕜, A)` is a normed algebra, and, when `A` is a C⋆-algebra, we show +that `𝓜(𝕜, A)` is also a C⋆-algebra. Moreover, in this case, for `a : 𝓜(𝕜, A)`, +`‖a‖ = ‖a.fst‖ = ‖a.snd‖`. -/ + +/-- The normed group structure is inherited as the pullback under the ring monomoprhism +`double_centralizer.to_prod_mul_opposite_hom : 𝓜(𝕜, A) →+* (A →L[𝕜] A) × (A →L[𝕜] A)ᵐᵒᵖ`. -/ +noncomputable instance : normed_ring 𝓜(𝕜, A) := +normed_ring.induced _ _ (to_prod_mul_opposite_hom : 𝓜(𝕜, A) →+* (A →L[𝕜] A) × (A →L[𝕜] A)ᵐᵒᵖ) + to_prod_mul_opposite_injective + +-- even though the definition is actually in terms of `double_centralizer.to_prod_mul_opposite`, we +-- choose to see through that here to avoid `mul_opposite.op` appearing. +lemma norm_def (a : 𝓜(𝕜, A)) : ‖a‖ = ‖a.to_prod_hom‖ := rfl +lemma nnnorm_def (a : 𝓜(𝕜, A)) : ‖a‖₊ = ‖a.to_prod_hom‖₊ := rfl + +lemma norm_def' (a : 𝓜(𝕜, A)) : ‖a‖ = ‖a.to_prod_mul_opposite_hom‖ := rfl +lemma nnnorm_def' (a : 𝓜(𝕜, A)) : ‖a‖₊ = ‖a.to_prod_mul_opposite_hom‖₊ := rfl + +instance : normed_space 𝕜 𝓜(𝕜, A) := +{ norm_smul_le := λ k a, (norm_smul k a.to_prod_mul_opposite).le, + .. double_centralizer.module } + +instance : normed_algebra 𝕜 𝓜(𝕜, A) := +{ ..double_centralizer.algebra, ..double_centralizer.normed_space } + +lemma uniform_embedding_to_prod_mul_opposite : + uniform_embedding (@to_prod_mul_opposite 𝕜 A _ _ _ _ _) := +uniform_embedding_comap to_prod_mul_opposite_injective + +instance [complete_space A] : complete_space 𝓜(𝕜, A) := +begin + rw complete_space_iff_is_complete_range + uniform_embedding_to_prod_mul_opposite.to_uniform_inducing, + apply is_closed.is_complete, + simp only [range_to_prod_mul_opposite, set.set_of_forall], + refine is_closed_Inter (λ x, is_closed_Inter $ λ y, is_closed_eq _ _), + exact ((continuous_linear_map.apply 𝕜 A _).continuous.comp $ + continuous_unop.comp continuous_snd).mul continuous_const, + exact continuous_const.mul ((continuous_linear_map.apply 𝕜 A _).continuous.comp continuous_fst), +end + +variables [star_ring A] [cstar_ring A] + +/-- For `a : 𝓜(𝕜, A)`, the norms of `a.fst` and `a.snd` coincide, and hence these +also coincide with `‖a‖` which is `max (‖a.fst‖) (‖a.snd‖)`. -/ +lemma norm_fst_eq_snd (a : 𝓜(𝕜, A)) : ‖a.fst‖ = ‖a.snd‖ := +begin + -- a handy lemma for this proof + have h0 : ∀ f : A →L[𝕜] A, ∀ C : ℝ≥0, (∀ b : A, ‖f b‖₊ ^ 2 ≤ C * ‖f b‖₊ * ‖b‖₊) → ‖f‖₊ ≤ C, + { intros f C h, + have h1 : ∀ b, C * ‖f b‖₊ * ‖b‖₊ ≤ C * ‖f‖₊ * ‖b‖₊ ^ 2, + { intros b, + convert mul_le_mul_right' (mul_le_mul_left' (f.le_op_nnnorm b) C) (‖b‖₊) using 1, + ring, }, + have := div_le_of_le_mul (f.op_nnnorm_le_bound _ (by simpa only [sqrt_sq, sqrt_mul] + using (λ b, sqrt_le_sqrt_iff.mpr ((h b).trans (h1 b))))), + convert rpow_le_rpow this two_pos.le, + { simp only [rpow_two, div_pow, sq_sqrt], simp only [sq, mul_self_div_self] }, + { simp only [rpow_two, sq_sqrt] } }, + have h1 : ∀ b, ‖a.fst b‖₊ ^ 2 ≤ ‖a.snd‖₊ * ‖a.fst b‖₊ * ‖b‖₊, + { intros b, + calc ‖a.fst b‖₊ ^ 2 + = ‖star (a.fst b) * (a.fst b)‖₊ + : by simpa only [←sq] using (cstar_ring.nnnorm_star_mul_self).symm + ... ≤ ‖a.snd (star (a.fst b))‖₊ * ‖b‖₊ : a.central (star (a.fst b)) b ▸ nnnorm_mul_le _ _ + ... ≤ ‖a.snd‖₊ * ‖a.fst b‖₊ * ‖b‖₊ + : nnnorm_star (a.fst b) ▸ mul_le_mul_right' (a.snd.le_op_nnnorm _) _}, + have h2 : ∀ b, ‖a.snd b‖₊ ^ 2 ≤ ‖a.fst‖₊ * ‖a.snd b‖₊ * ‖b‖₊, + { intros b, + calc ‖a.snd b‖₊ ^ 2 + = ‖a.snd b * star (a.snd b)‖₊ + : by simpa only [←sq] using (cstar_ring.nnnorm_self_mul_star).symm + ... ≤ ‖b‖₊ * ‖a.fst (star (a.snd b))‖₊ + : (a.central b (star (a.snd b))).symm ▸ nnnorm_mul_le _ _ + ... = ‖a.fst (star (a.snd b))‖₊ * ‖b‖₊ : mul_comm _ _ + ... ≤ ‖a.fst‖₊ * ‖a.snd b‖₊ * ‖b‖₊ + : nnnorm_star (a.snd b) ▸ mul_le_mul_right' (a.fst.le_op_nnnorm _) _ }, + exact le_antisymm (h0 _ _ h1) (h0 _ _ h2), +end + +lemma nnnorm_fst_eq_snd (a : 𝓜(𝕜, A)) : ‖a.fst‖₊ = ‖a.snd‖₊ := subtype.ext $ norm_fst_eq_snd a +@[simp] lemma norm_fst (a : 𝓜(𝕜, A)) : ‖a.fst‖ = ‖a‖ := + by simp only [norm_def, to_prod_hom_apply, prod.norm_def, norm_fst_eq_snd, max_eq_right, + eq_self_iff_true] +@[simp] lemma norm_snd (a : 𝓜(𝕜, A)) : ‖a.snd‖ = ‖a‖ := by rw [←norm_fst, norm_fst_eq_snd] +@[simp] lemma nnnorm_fst (a : 𝓜(𝕜, A)) : ‖a.fst‖₊ = ‖a‖₊ := subtype.ext (norm_fst a) +@[simp] lemma nnnorm_snd (a : 𝓜(𝕜, A)) : ‖a.snd‖₊ = ‖a‖₊ := subtype.ext (norm_snd a) + +end nontrivially_normed + +section densely_normed + +variables {𝕜 A : Type*} [densely_normed_field 𝕜] [star_ring 𝕜] +variables [non_unital_normed_ring A] [star_ring A] [cstar_ring A] +variables [normed_space 𝕜 A] [smul_comm_class 𝕜 A A] [is_scalar_tower 𝕜 A A] [star_module 𝕜 A] + +instance : cstar_ring 𝓜(𝕜, A) := +{ norm_star_mul_self := λ a, congr_arg (coe : ℝ≥0 → ℝ) $ show ‖star a * a‖₊ = ‖a‖₊ * ‖a‖₊, from + begin + /- The essence of the argument is this: let `a = (L,R)` and recall `‖a‖ = ‖L‖`. + `star a = (star ∘ R ∘ star, star ∘ L ∘ star)`. Then for any `x y : A`, we have + `‖star a * a‖ = ‖(star a * a).snd‖ = ‖R (star (L (star x))) * y‖ = ‖star (L (star x)) * L y‖` + Now, on the one hand, + `‖star (L (star x)) * L y‖ ≤ ‖star (L (star x))‖ * ‖L y‖ = ‖L (star x)‖ * ‖L y‖ ≤ ‖L‖ ^ 2` + whenever `‖x‖, ‖y‖ ≤ 1`, so the supremum over all such `x, y` is at most `‖L‖ ^ 2`. + On the other hand, for any `‖z‖ ≤ 1`, we may choose `x := star z` and `y := z` to get: + `‖star (L (star x)) * L y‖ = ‖star (L z) * (L z)‖ = ‖L z‖ ^ 2`, and taking the supremum over + all such `z` yields that the supremum is at least `‖L‖ ^ 2`. It is the latter part of the + argument where `densely_normed_field 𝕜` is required (for `Sup_closed_unit_ball_eq_nnnorm`). -/ + have hball : (metric.closed_ball (0 : A) 1).nonempty := + metric.nonempty_closed_ball.2 (zero_le_one), + have key : ∀ x y, ‖x‖₊ ≤ 1 → ‖y‖₊ ≤ 1 → ‖a.snd (star (a.fst (star x))) * y‖₊ ≤ ‖a‖₊ * ‖a‖₊, + { intros x y hx hy, + rw [a.central], + calc ‖star (a.fst (star x)) * a.fst y‖₊ ≤ ‖a.fst (star x)‖₊ * ‖a.fst y‖₊ + : nnnorm_star (a.fst (star x)) ▸ nnnorm_mul_le _ _ + ... ≤ (‖a.fst‖₊ * 1) * (‖a.fst‖₊ * 1) + : mul_le_mul' (a.fst.le_op_norm_of_le ((nnnorm_star x).trans_le hx)) + (a.fst.le_op_norm_of_le hy) + ... ≤ ‖a‖₊ * ‖a‖₊ : by simp only [mul_one, nnnorm_fst] }, + rw ←nnnorm_snd, + simp only [mul_snd, ←Sup_closed_unit_ball_eq_nnnorm, star_snd, mul_apply], + simp only [←@op_nnnorm_mul 𝕜 A], + simp only [←Sup_closed_unit_ball_eq_nnnorm, mul_apply'], + refine cSup_eq_of_forall_le_of_forall_lt_exists_gt (hball.image _) _ (λ r hr, _), + { rintro - ⟨x, hx, rfl⟩, + refine cSup_le (hball.image _) _, + rintro - ⟨y, hy, rfl⟩, + exact key x y (mem_closed_ball_zero_iff.1 hx) (mem_closed_ball_zero_iff.1 hy) }, + { simp only [set.mem_image, set.mem_set_of_eq, exists_prop, exists_exists_and_eq_and], + have hr' : r.sqrt < ‖a‖₊ := (‖a‖₊).sqrt_mul_self ▸ nnreal.sqrt_lt_sqrt_iff.2 hr, + simp_rw [←nnnorm_fst, ←Sup_closed_unit_ball_eq_nnnorm] at hr', + obtain ⟨_, ⟨x, hx, rfl⟩, hxr⟩ := exists_lt_of_lt_cSup (hball.image _) hr', + have hx' : ‖x‖₊ ≤ 1 := mem_closed_ball_zero_iff.1 hx, + refine ⟨star x, mem_closed_ball_zero_iff.2 ((nnnorm_star x).trans_le hx'), _⟩, + refine lt_cSup_of_lt _ ⟨x, hx, rfl⟩ _, + { refine ⟨‖a‖₊ * ‖a‖₊, _⟩, + rintros - ⟨y, hy, rfl⟩, + exact key (star x) y ((nnnorm_star x).trans_le hx') (mem_closed_ball_zero_iff.1 hy) }, + { simpa only [a.central, star_star, cstar_ring.nnnorm_star_mul_self, nnreal.sq_sqrt, ←sq] + using pow_lt_pow_of_lt_left hxr zero_le' two_pos } } + end } + +end densely_normed + +end double_centralizer From ef7acf407d265ad4081c8998687e994fa80ba70c Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Fri, 24 Feb 2023 07:37:37 +0000 Subject: [PATCH 0517/1291] chore(*): add mathlib4 synchronization comments (#18490) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `category_theory.bicategory.End` * `category_theory.essentially_small` * `category_theory.functor.epi_mono` * `category_theory.functor.reflects_isomorphisms` * `category_theory.monoidal.functor` * `data.qpf.multivariate.constructions.fix` * `dynamics.flow` * `group_theory.noncomm_pi_coprod` * `linear_algebra.linear_independent` * `set_theory.cardinal.ordinal` --- src/category_theory/bicategory/End.lean | 3 +++ src/category_theory/essentially_small.lean | 3 +++ src/category_theory/functor/epi_mono.lean | 3 +++ src/category_theory/functor/reflects_isomorphisms.lean | 3 +++ src/category_theory/monoidal/functor.lean | 3 +++ src/data/qpf/multivariate/constructions/fix.lean | 3 +++ src/dynamics/flow.lean | 3 +++ src/group_theory/noncomm_pi_coprod.lean | 3 +++ src/linear_algebra/linear_independent.lean | 3 +++ src/set_theory/cardinal/ordinal.lean | 3 +++ 10 files changed, 30 insertions(+) diff --git a/src/category_theory/bicategory/End.lean b/src/category_theory/bicategory/End.lean index f677e134b87f8..a5c38f929d6eb 100644 --- a/src/category_theory/bicategory/End.lean +++ b/src/category_theory/bicategory/End.lean @@ -8,6 +8,9 @@ import category_theory.monoidal.category /-! # Endomorphisms of an object in a bicategory, as a monoidal category. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ namespace category_theory diff --git a/src/category_theory/essentially_small.lean b/src/category_theory/essentially_small.lean index df2db56995524..cb4fa7f70fd77 100644 --- a/src/category_theory/essentially_small.lean +++ b/src/category_theory/essentially_small.lean @@ -10,6 +10,9 @@ import category_theory.skeletal /-! # Essentially small categories. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A category given by `(C : Type u) [category.{v} C]` is `w`-essentially small if there exists a `small_model C : Type w` equipped with `[small_category (small_model C)]`. diff --git a/src/category_theory/functor/epi_mono.lean b/src/category_theory/functor/epi_mono.lean index 521a68d8ff733..2a07f378344ee 100644 --- a/src/category_theory/functor/epi_mono.lean +++ b/src/category_theory/functor/epi_mono.lean @@ -10,6 +10,9 @@ import category_theory.lifting_properties.adjunction /-! # Preservation and reflection of monomorphisms and epimorphisms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We provide typeclasses that state that a functor preserves or reflects monomorphisms or epimorphisms. -/ diff --git a/src/category_theory/functor/reflects_isomorphisms.lean b/src/category_theory/functor/reflects_isomorphisms.lean index f017a78d62693..ce74ef7f9e527 100644 --- a/src/category_theory/functor/reflects_isomorphisms.lean +++ b/src/category_theory/functor/reflects_isomorphisms.lean @@ -10,6 +10,9 @@ import category_theory.functor.fully_faithful /-! # Functors which reflect isomorphisms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A functor `F` reflects isomorphisms if whenever `F.map f` is an isomorphism, `f` was too. It is formalized as a `Prop` valued typeclass `reflects_isomorphisms F`. diff --git a/src/category_theory/monoidal/functor.lean b/src/category_theory/monoidal/functor.lean index d43a48e845e3c..18f74762dcea1 100644 --- a/src/category_theory/monoidal/functor.lean +++ b/src/category_theory/monoidal/functor.lean @@ -10,6 +10,9 @@ import category_theory.products.basic /-! # (Lax) monoidal functors +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A lax monoidal functor `F` between monoidal categories `C` and `D` is a functor between the underlying categories equipped with morphisms * `ε : 𝟙_ D ⟶ F.obj (𝟙_ C)` (called the unit morphism) diff --git a/src/data/qpf/multivariate/constructions/fix.lean b/src/data/qpf/multivariate/constructions/fix.lean index ef1470eae2ebb..2c1c33ceed19e 100644 --- a/src/data/qpf/multivariate/constructions/fix.lean +++ b/src/data/qpf/multivariate/constructions/fix.lean @@ -9,6 +9,9 @@ import data.qpf.multivariate.basic /-! # The initial algebra of a multivariate qpf is again a qpf. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + For a `(n+1)`-ary QPF `F (α₀,..,αₙ)`, we take the least fixed point of `F` with regards to its last argument `αₙ`. The result is a `n`-ary functor: `fix F (α₀,..,αₙ₋₁)`. Making `fix F` into a functor allows us to take the fixed point, compose with other functors diff --git a/src/dynamics/flow.lean b/src/dynamics/flow.lean index 3235f5a1b40f2..14000867f74f3 100644 --- a/src/dynamics/flow.lean +++ b/src/dynamics/flow.lean @@ -10,6 +10,9 @@ import logic.function.iterate /-! # Flows and invariant sets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines a flow on a topological space `α` by a topological monoid `τ` as a continuous monoid-act of `τ` on `α`. Anticipating the cases where `τ` is one of `ℕ`, `ℤ`, `ℝ⁺`, or `ℝ`, we use additive diff --git a/src/group_theory/noncomm_pi_coprod.lean b/src/group_theory/noncomm_pi_coprod.lean index 3e2db788d60d9..602cb46e88539 100644 --- a/src/group_theory/noncomm_pi_coprod.lean +++ b/src/group_theory/noncomm_pi_coprod.lean @@ -12,6 +12,9 @@ import order.sup_indep /-! # Canonical homomorphism from a finite family of monoids +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the construction of the canonical homomorphism from a family of monoids. Given a family of morphisms `ϕ i : N i →* M` for each `i : ι` where elements in the diff --git a/src/linear_algebra/linear_independent.lean b/src/linear_algebra/linear_independent.lean index ce81c6baf68e4..6a7e065316d26 100644 --- a/src/linear_algebra/linear_independent.lean +++ b/src/linear_algebra/linear_independent.lean @@ -12,6 +12,9 @@ import set_theory.cardinal.basic # Linear independence +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines linear independence in a module or vector space. It is inspired by Isabelle/HOL's linear algebra, and hence indirectly by HOL Light. diff --git a/src/set_theory/cardinal/ordinal.lean b/src/set_theory/cardinal/ordinal.lean index 208a28970e325..5ffa366bf8606 100644 --- a/src/set_theory/cardinal/ordinal.lean +++ b/src/set_theory/cardinal/ordinal.lean @@ -12,6 +12,9 @@ import tactic.linarith /-! # Cardinals and ordinals +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Relationships between cardinals and ordinals, properties of cardinals that are proved using ordinals. From 8e57ff9a1d998da9ab3da9256bada5ba51a35a72 Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Fri, 24 Feb 2023 10:27:56 +0000 Subject: [PATCH 0518/1291] feat(geometry/manifold/cont_mdiff): local structomorphisms of `cont_diff_groupoid` (#17291) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Let `M` and `M'` be smooth manifolds with the same model-with-corners, `I`. Then `f : M → M'` is a local structomorphism for `I`, if and only if it is manifold-smooth on the domain of definition in both directions. Co-authored-by: Floris van Doorn --- src/geometry/manifold/cont_mdiff.lean | 160 +++++++++++++++++- .../manifold/local_invariant_properties.lean | 58 +++++++ 2 files changed, 216 insertions(+), 2 deletions(-) diff --git a/src/geometry/manifold/cont_mdiff.lean b/src/geometry/manifold/cont_mdiff.lean index 66382c546e2e2..7538d4211da58 100644 --- a/src/geometry/manifold/cont_mdiff.lean +++ b/src/geometry/manifold/cont_mdiff.lean @@ -158,10 +158,10 @@ begin end lemma cont_diff_within_at_prop_id (x : H) : - cont_diff_within_at_prop I I ∞ id univ x := + cont_diff_within_at_prop I I n id univ x := begin simp [cont_diff_within_at_prop], - have : cont_diff_within_at 𝕜 ∞ id (range I) (I x) := + have : cont_diff_within_at 𝕜 n id (range I) (I x) := cont_diff_id.cont_diff_at.cont_diff_within_at, apply this.congr (λ y hy, _), { simp only with mfld_simps }, @@ -1089,6 +1089,14 @@ lemma cont_mdiff_on_ext_chart_at : cont_mdiff_on I 𝓘(𝕜, E) n (ext_chart_at I x) (chart_at H x).source := λ x' hx', (cont_mdiff_at_ext_chart_at' hx').cont_mdiff_within_at +omit Is + +/-- An element of `cont_diff_groupoid ⊤ I` is `C^n` for any `n`. -/ +lemma cont_mdiff_on_of_mem_cont_diff_groupoid {e' : local_homeomorph H H} + (h : e' ∈ cont_diff_groupoid ⊤ I) : cont_mdiff_on I I n e' e'.source := +(cont_diff_within_at_local_invariant_prop I I n).lift_prop_on_of_mem_groupoid + (cont_diff_within_at_prop_id I) h + end atlas /-! ### The identity is smooth -/ @@ -1676,3 +1684,151 @@ hf.smul hg lemma smooth.smul {f : M → 𝕜} {g : M → V} (hf : smooth I 𝓘(𝕜) f) (hg : smooth I 𝓘(𝕜, V) g) : smooth I 𝓘(𝕜, V) (λ p, f p • g p) := hf.smul hg + +/-! ### Smoothness of (local) structomorphisms -/ +section + +variables [charted_space H M'] [IsM' : smooth_manifold_with_corners I M'] +include Is IsM' + +lemma is_local_structomorph_on_cont_diff_groupoid_iff_aux {f : local_homeomorph M M'} + (hf : lift_prop_on (cont_diff_groupoid ⊤ I).is_local_structomorph_within_at f f.source) : + smooth_on I I f f.source := +begin + -- It suffices to show smoothness near each `x` + apply cont_mdiff_on_of_locally_cont_mdiff_on, + intros x hx, + let c := chart_at H x, + let c' := chart_at H (f x), + obtain ⟨-, hxf⟩ := hf x hx, + -- Since `f` is a local structomorph, it is locally equal to some transferred element `e` of + -- the `cont_diff_groupoid`. + obtain ⟨e, he, he' : eq_on (c' ∘ f ∘ c.symm) e (c.symm ⁻¹' f.source ∩ e.source), + hex : c x ∈ e.source⟩ := hxf (by simp only [hx] with mfld_simps), + -- We choose a convenient set `s` in `M`. + let s : set M := (f.trans c').source ∩ ((c.trans e).trans c'.symm).source, + refine ⟨s, (f.trans c').open_source.inter ((c.trans e).trans c'.symm).open_source, _, _⟩, + { simp only with mfld_simps, + rw ← he'; simp only [hx, hex] with mfld_simps }, + -- We need to show `f` is `cont_mdiff_on` the domain `s ∩ f.source`. We show this in two + -- steps: `f` is equal to `c'.symm ∘ e ∘ c` on that domain and that function is + -- `cont_mdiff_on` it. + have H₁ : cont_mdiff_on I I ⊤ (c'.symm ∘ e ∘ c) s, + { have hc' : cont_mdiff_on I I ⊤ c'.symm _ := cont_mdiff_on_chart_symm, + have he'' : cont_mdiff_on I I ⊤ e _ := cont_mdiff_on_of_mem_cont_diff_groupoid he, + have hc : cont_mdiff_on I I ⊤ c _ := cont_mdiff_on_chart, + refine (hc'.comp' (he''.comp' hc)).mono _, + mfld_set_tac }, + have H₂ : eq_on f (c'.symm ∘ e ∘ c) s, + { intros y hy, + simp only with mfld_simps at hy, + have hy₁ : f y ∈ c'.source := by simp only [hy] with mfld_simps, + have hy₂ : y ∈ c.source := by simp only [hy] with mfld_simps, + have hy₃ : c y ∈ c.symm ⁻¹' f.source ∩ e.source := by simp only [hy] with mfld_simps, + calc f y = c'.symm (c' (f y)) : by rw c'.left_inv hy₁ + ... = c'.symm (c' (f (c.symm (c y)))) : by rw c.left_inv hy₂ + ... = c'.symm (e (c y)) : by rw ← he' hy₃ }, + refine (H₁.congr H₂).mono _, + mfld_set_tac +end + +/-- Let `M` and `M'` be smooth manifolds with the same model-with-corners, `I`. Then `f : M → M'` +is a local structomorphism for `I`, if and only if it is manifold-smooth on the domain of definition +in both directions. -/ +lemma is_local_structomorph_on_cont_diff_groupoid_iff (f : local_homeomorph M M') : + lift_prop_on (cont_diff_groupoid ⊤ I).is_local_structomorph_within_at f f.source + ↔ smooth_on I I f f.source ∧ smooth_on I I f.symm f.target := +begin + split, + { intros h, + refine ⟨is_local_structomorph_on_cont_diff_groupoid_iff_aux h, + is_local_structomorph_on_cont_diff_groupoid_iff_aux _⟩, + -- todo: we can generalize this part of the proof to a lemma + intros X hX, + let x := f.symm X, + have hx : x ∈ f.source := f.symm.maps_to hX, + let c := chart_at H x, + let c' := chart_at H X, + obtain ⟨-, hxf⟩ := h x hx, + refine ⟨(f.symm.continuous_at hX).continuous_within_at, λ h2x, _⟩, + obtain ⟨e, he, h2e, hef, hex⟩ : ∃ e : local_homeomorph H H, e ∈ cont_diff_groupoid ⊤ I ∧ + e.source ⊆ (c.symm ≫ₕ f ≫ₕ c').source ∧ + eq_on (c' ∘ f ∘ c.symm) e e.source ∧ c x ∈ e.source, + { have h1 : c' = chart_at H (f x) := by simp only [f.right_inv hX], + have h2 : ⇑c' ∘ ⇑f ∘ ⇑(c.symm) = ⇑(c.symm ≫ₕ f ≫ₕ c') := rfl, + have hcx : c x ∈ c.symm ⁻¹' f.source, { simp only [hx] with mfld_simps }, + rw [h2], + rw [← h1, h2, local_homeomorph.is_local_structomorph_within_at_iff'] at hxf, + { exact hxf hcx }, + { mfld_set_tac }, + { apply or.inl, + simp only [hx, h1] with mfld_simps } }, + have h2X : c' X = e (c (f.symm X)), + { rw ← hef hex, + dsimp only [function.comp], + have hfX : f.symm X ∈ c.source := by simp only [hX] with mfld_simps, + rw [c.left_inv hfX, f.right_inv hX] }, + have h3e : eq_on (c ∘ f.symm ∘ c'.symm) e.symm (c'.symm ⁻¹' f.target ∩ e.target), + { have h1 : eq_on (c.symm ≫ₕ f ≫ₕ c').symm e.symm (e.target ∩ e.target), + { apply eq_on.symm, + refine e.is_image_source_target.symm_eq_on_of_inter_eq_of_eq_on _ _, + { rw [inter_self, inter_eq_right_iff_subset.mpr h2e] }, + rw [inter_self], exact hef.symm }, + have h2 : e.target ⊆ (c.symm ≫ₕ f ≫ₕ c').target, + { intros x hx, rw [← e.right_inv hx, ← hef (e.symm.maps_to hx)], + exact local_homeomorph.maps_to _ (h2e $ e.symm.maps_to hx) }, + rw [inter_self] at h1, + rwa [inter_eq_right_iff_subset.mpr], + refine h2.trans _, + mfld_set_tac }, + refine ⟨e.symm, structure_groupoid.symm _ he, h3e, _⟩, + rw [h2X], exact e.maps_to hex }, + { -- We now show the converse: a local homeomorphism `f : M → M'` which is smooth in both + -- directions is a local structomorphism. We do this by proposing + -- `((chart_at H x).symm.trans f).trans (chart_at H (f x))` as a candidate for a structomorphism + -- of `H`. + rintros ⟨h₁, h₂⟩ x hx, + refine ⟨(h₁ x hx).continuous_within_at, _⟩, + let c := chart_at H x, + let c' := chart_at H (f x), + rintros (hx' : c x ∈ c.symm ⁻¹' f.source), + -- propose `(c.symm.trans f).trans c'` as a candidate for a local structomorphism of `H` + refine ⟨(c.symm.trans f).trans c', ⟨_, _⟩, (_ : eq_on (c' ∘ f ∘ c.symm) _ _), _⟩, + { -- smoothness of the candidate local structomorphism in the forward direction + intros y hy, + simp only with mfld_simps at hy, + have H : cont_mdiff_within_at I I ⊤ f (f ≫ₕ c').source (((ext_chart_at I x).symm) y), + { refine (h₁ ((ext_chart_at I x).symm y) _).mono _, + { simp only [hy] with mfld_simps }, + { mfld_set_tac } }, + have hy' : (ext_chart_at I x).symm y ∈ c.source := by simp only [hy] with mfld_simps, + have hy'' : f ((ext_chart_at I x).symm y) ∈ c'.source := by simp only [hy] with mfld_simps, + rw cont_mdiff_within_at_iff_of_mem_source hy' hy'' at H, + { convert H.2.mono _, + { simp only [hy] with mfld_simps }, + { mfld_set_tac } }, + { apply_instance }, + { apply_instance } }, + { -- smoothness of the candidate local structomorphism in the reverse direction + intros y hy, + simp only with mfld_simps at hy, + have H : cont_mdiff_within_at I I ⊤ f.symm (f.symm ≫ₕ c).source + (((ext_chart_at I (f x)).symm) y), + { refine (h₂ ((ext_chart_at I (f x)).symm y) _).mono _, + { simp only [hy] with mfld_simps }, + { mfld_set_tac } }, + have hy' : (ext_chart_at I (f x)).symm y ∈ c'.source := by simp only [hy] with mfld_simps, + have hy'' : f.symm ((ext_chart_at I (f x)).symm y) ∈ c.source, + { simp only [hy] with mfld_simps }, + rw cont_mdiff_within_at_iff_of_mem_source hy' hy'' at H, + { convert H.2.mono _, + { simp only [hy] with mfld_simps }, + { mfld_set_tac } }, + { apply_instance }, + { apply_instance } }, + -- now check the candidate local structomorphism agrees with `f` where it is supposed to + { simp only with mfld_simps }, + { simp only [hx'] with mfld_simps } }, +end + +end diff --git a/src/geometry/manifold/local_invariant_properties.lean b/src/geometry/manifold/local_invariant_properties.lean index 2ba32fbfda0dd..e785862cb25f2 100644 --- a/src/geometry/manifold/local_invariant_properties.lean +++ b/src/geometry/manifold/local_invariant_properties.lean @@ -549,6 +549,16 @@ lemma lift_prop_on_chart_symm [has_groupoid M G] lift_prop_on Q (chart_at H x).symm (chart_at H x).target := hG.lift_prop_on_symm_of_mem_maximal_atlas hQ (chart_mem_maximal_atlas G x) +lemma lift_prop_at_of_mem_groupoid (hG : G.local_invariant_prop G Q) (hQ : ∀ y, Q id univ y) + {f : local_homeomorph H H} (hf : f ∈ G) {x : H} (hx : x ∈ f.source) : + lift_prop_at Q f x := +lift_prop_at_of_mem_maximal_atlas hG hQ (G.mem_maximal_atlas_of_mem_groupoid hf) hx + +lemma lift_prop_on_of_mem_groupoid (hG : G.local_invariant_prop G Q) (hQ : ∀ y, Q id univ y) + {f : local_homeomorph H H} (hf : f ∈ G) : + lift_prop_on Q f f.source := +lift_prop_on_of_mem_maximal_atlas hG hQ (G.mem_maximal_atlas_of_mem_groupoid hf) + lemma lift_prop_id (hG : G.local_invariant_prop G Q) (hQ : ∀ y, Q id univ y) : lift_prop Q (id : M → M) := begin @@ -615,6 +625,54 @@ lemma is_local_structomorph_within_at_local_invariant_prop [closed_under_restric { simpa only [hex, hef ⟨hx, hex⟩] with mfld_simps using hfx } end } +/-- A slight reformulation of `is_local_structomorph_within_at` when `f` is a local homeomorph. + This gives us an `e` that is defined on a subset of `f.source`. -/ +lemma _root_.local_homeomorph.is_local_structomorph_within_at_iff {G : structure_groupoid H} + [closed_under_restriction G] + (f : local_homeomorph H H) {s : set H} {x : H} (hx : x ∈ f.source ∪ sᶜ) : + G.is_local_structomorph_within_at ⇑f s x ↔ + x ∈ s → ∃ (e : local_homeomorph H H), e ∈ G ∧ e.source ⊆ f.source ∧ + eq_on f ⇑e (s ∩ e.source) ∧ x ∈ e.source := +begin + split, + { intros hf h2x, + obtain ⟨e, he, hfe, hxe⟩ := hf h2x, + refine ⟨e.restr f.source, closed_under_restriction' he f.open_source, _, _, hxe, _⟩, + { simp_rw [local_homeomorph.restr_source], + refine (inter_subset_right _ _).trans interior_subset }, + { intros x' hx', exact hfe ⟨hx'.1, hx'.2.1⟩ }, + { rw [f.open_source.interior_eq], exact or.resolve_right hx (not_not.mpr h2x) } }, + { intros hf hx, obtain ⟨e, he, h2e, hfe, hxe⟩ := hf hx, exact ⟨e, he, hfe, hxe⟩ } +end + +/-- A slight reformulation of `is_local_structomorph_within_at` when `f` is a local homeomorph and + the set we're considering is a superset of `f.source`. -/ +lemma _root_.local_homeomorph.is_local_structomorph_within_at_iff' {G : structure_groupoid H} + [closed_under_restriction G] + (f : local_homeomorph H H) {s : set H} {x : H} (hs : f.source ⊆ s) (hx : x ∈ f.source ∪ sᶜ) : + G.is_local_structomorph_within_at ⇑f s x ↔ + x ∈ s → ∃ (e : local_homeomorph H H), e ∈ G ∧ e.source ⊆ f.source ∧ + eq_on f ⇑e e.source ∧ x ∈ e.source := +begin + simp_rw [f.is_local_structomorph_within_at_iff hx], + refine imp_congr_right (λ hx, exists_congr $ λ e, and_congr_right $ λ he, _), + refine and_congr_right (λ h2e, _), + rw [inter_eq_right_iff_subset.mpr (h2e.trans hs)], +end + +/-- A slight reformulation of `is_local_structomorph_within_at` when `f` is a local homeomorph and + the set we're considering is `f.source`. -/ +lemma _root_.local_homeomorph.is_local_structomorph_within_at_source_iff {G : structure_groupoid H} + [closed_under_restriction G] + (f : local_homeomorph H H) {x : H} : + G.is_local_structomorph_within_at ⇑f f.source x ↔ + x ∈ f.source → ∃ (e : local_homeomorph H H), e ∈ G ∧ e.source ⊆ f.source ∧ + eq_on f ⇑e e.source ∧ x ∈ e.source := +begin + have : x ∈ f.source ∪ f.sourceᶜ, { simp_rw [union_compl_self] }, + exact f.is_local_structomorph_within_at_iff' subset.rfl this, +end + variables {H₁ : Type*} [topological_space H₁] {H₂ : Type*} [topological_space H₂] {H₃ : Type*} [topological_space H₃] [charted_space H₁ H₂] [charted_space H₂ H₃] {G₁ : structure_groupoid H₁} [has_groupoid H₂ G₁] [closed_under_restriction G₁] From 634284e765db7adb83fccbc9ad5fe42c0f2a8bb3 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 24 Feb 2023 10:27:58 +0000 Subject: [PATCH 0519/1291] refactor(data/real/ereal): add a two-argument induction principle (#18481) We use exactly the same induction process whenever proving anything about `*` on `ereal`, so we may as well package it into a helper lemma. Co-authored-by: sgouezel --- src/data/real/ereal.lean | 238 ++++++++++++++++++++++----------------- 1 file changed, 134 insertions(+), 104 deletions(-) diff --git a/src/data/real/ereal.lean b/src/data/real/ereal.lean index a3a5bf5e7f728..7007030abf71c 100644 --- a/src/data/real/ereal.lean +++ b/src/data/real/ereal.lean @@ -63,6 +63,10 @@ def real.to_ereal : ℝ → ereal := some ∘ some namespace ereal +-- things unify with `with_bot.decidable_lt` later if we we don't provide this explicitly. +instance decidable_lt : decidable_rel ((<) : ereal → ereal → Prop) := +with_bot.decidable_lt + -- TODO: Provide explicitly, otherwise it is inferred noncomputably from `complete_linear_order` instance : has_top ereal := ⟨some ⊤⟩ @@ -121,6 +125,42 @@ protected def mul : ereal → ereal → ereal instance : has_mul ereal := ⟨ereal.mul⟩ +/-- Induct on two ereals by performing case splits on the sign of one whenever the other is +infinite. -/ +@[elab_as_eliminator] +lemma induction₂ {P : ereal → ereal → Prop} + (top_top : P ⊤ ⊤) + (top_pos : ∀ x : ℝ, 0 < x → P ⊤ x) + (top_zero : P ⊤ 0) + (top_neg : ∀ x : ℝ, x < 0 → P ⊤ x) + (top_bot : P ⊤ ⊥) + (pos_top : ∀ x : ℝ, 0 < x → P x ⊤) + (pos_bot : ∀ x : ℝ, 0 < x → P x ⊥) + (zero_top : P 0 ⊤) + (coe_coe : ∀ x y : ℝ, P x y) + (zero_bot : P 0 ⊥) + (neg_top : ∀ x : ℝ, x < 0 → P x ⊤) + (neg_bot : ∀ x : ℝ, x < 0 → P x ⊥) + (bot_top : P ⊥ ⊤) + (bot_pos : ∀ x : ℝ, 0 < x → P ⊥ x) + (bot_zero : P ⊥ 0) + (bot_neg : ∀ x : ℝ, x < 0 → P ⊥ x) + (bot_bot : P ⊥ ⊥) : + ∀ x y, P x y +| ⊥ ⊥ := bot_bot +| ⊥ (y : ℝ) := + by { rcases lt_trichotomy 0 y with hy|rfl|hy, exacts [bot_pos y hy, bot_zero, bot_neg y hy] } +| ⊥ ⊤ := bot_top +| (x : ℝ) ⊥ := + by { rcases lt_trichotomy 0 x with hx|rfl|hx, exacts [pos_bot x hx, zero_bot, neg_bot x hx] } +| (x : ℝ) (y : ℝ) := coe_coe _ _ +| (x : ℝ) ⊤ := + by { rcases lt_trichotomy 0 x with hx|rfl|hx, exacts [pos_top x hx, zero_top, neg_top x hx] } +| ⊤ ⊥ := top_bot +| ⊤ (y : ℝ) := + by { rcases lt_trichotomy 0 y with hy|rfl|hy, exacts [top_pos y hy, top_zero, top_neg y hy] } +| ⊤ ⊤ := top_top + /-! `ereal` with its multiplication is a `comm_monoid_with_zero`. However, the proof of associativity by hand is extremely painful (with 125 cases...). Instead, we will deduce it later on from the facts that the absolute value and the sign are multiplicative functions taking value @@ -695,54 +735,57 @@ bot_mul_of_neg (ereal.coe_neg'.2 h) lemma to_real_mul {x y : ereal} : to_real (x * y) = to_real x * to_real y := begin - induction x using ereal.rec; induction y using ereal.rec, - { simp only [bot_mul_bot, to_real_top, to_real_bot, mul_zero] }, - { rcases lt_trichotomy 0 y with hy|rfl|hy, - { simp only [bot_mul_coe_of_pos hy, to_real_bot, zero_mul] }, - { simp only [coe_zero, mul_zero, to_real_zero] }, - { simp only [bot_mul_coe_of_neg hy, to_real_top, to_real_bot, zero_mul] } }, - { simp only [bot_mul_top, to_real_bot, to_real_top, mul_zero] }, - { rcases lt_trichotomy 0 x with hx|rfl|hx, - { simp only [coe_mul_bot_of_pos hx, to_real_bot, mul_zero] }, - { simp only [coe_zero, zero_mul, to_real_zero] }, - { simp only [coe_mul_bot_of_neg hx, to_real_top, to_real_bot, mul_zero] } }, - { simp only [← coe_mul, to_real_coe] }, - { rcases lt_trichotomy 0 x with hx|rfl|hx, - { simp only [coe_mul_top_of_pos hx, to_real_top, mul_zero] }, - { simp only [coe_zero, zero_mul, to_real_zero] }, - { simp only [coe_mul_top_of_neg hx, to_real_top, to_real_bot, mul_zero] } }, - { simp only [top_mul_bot, to_real_bot, mul_zero] }, - { rcases lt_trichotomy 0 y with hy|rfl|hy, - { simp only [top_mul_coe_of_pos hy, to_real_top, zero_mul] }, - { simp only [coe_zero, mul_zero, to_real_zero] }, - { simp only [top_mul_coe_of_neg hy, to_real_top, to_real_bot, zero_mul] } }, - { simp only [top_mul_top, to_real_top, mul_zero] } + -- TODO: replace with `induction using` in Lean 4, which supports multiple premises + with_cases + { apply @induction₂ (λ x y, to_real (x * y) = to_real x * to_real y) }; + propagate_tags { try { dsimp only} }, + case [top_zero, bot_zero, zero_top, zero_bot] { all_goals { simp only [zero_mul, mul_zero, + to_real_zero] } }, + case coe_coe : x y { norm_cast }, + case top_top { rw [top_mul_top, to_real_top, mul_zero] }, + case top_bot { rw [top_mul_bot, to_real_top, to_real_bot, zero_mul] }, + case bot_top { rw [bot_mul_top, to_real_bot, zero_mul] }, + case bot_bot { rw [bot_mul_bot, to_real_top, to_real_bot, zero_mul] }, + case pos_bot : x hx + { rw [to_real_bot, to_real_coe, coe_mul_bot_of_pos hx, to_real_bot, mul_zero] }, + case neg_bot : x hx + { rw [to_real_bot, to_real_coe, coe_mul_bot_of_neg hx, to_real_top, mul_zero] }, + case pos_top : x hx + { rw [to_real_top, to_real_coe, coe_mul_top_of_pos hx, to_real_top, mul_zero] }, + case neg_top : x hx + { rw [to_real_top, to_real_coe, coe_mul_top_of_neg hx, to_real_bot, mul_zero] }, + case top_pos : y hy + { rw [to_real_top, to_real_coe, top_mul_coe_of_pos hy, to_real_top, zero_mul] }, + case top_neg : y hy + { rw [to_real_top, to_real_coe, top_mul_coe_of_neg hy, to_real_bot, zero_mul] }, + case bot_pos : y hy + { rw [to_real_bot, to_real_coe, bot_mul_coe_of_pos hy, to_real_bot, zero_mul] }, + case bot_neg : y hy + { rw [to_real_bot, to_real_coe, bot_mul_coe_of_neg hy, to_real_top, zero_mul] }, end protected lemma neg_mul (x y : ereal) : -x * y = -(x * y) := begin - induction x using ereal.rec; induction y using ereal.rec, - { refl }, - { rcases lt_trichotomy 0 y with hy | rfl | hy, - { rw [bot_mul_coe_of_pos hy, neg_bot, top_mul_coe_of_pos hy] }, - { rw [coe_zero, mul_zero, mul_zero, neg_zero] }, - { rw [bot_mul_coe_of_neg hy, neg_bot, neg_top, top_mul_coe_of_neg hy] } }, - { refl }, - { rcases lt_trichotomy 0 x with hx | rfl | hx, - { rw [coe_mul_bot_of_pos hx, neg_bot, ← coe_neg, coe_mul_bot_of_neg (neg_neg_of_pos hx)], }, - { rw [coe_zero, zero_mul, neg_zero, zero_mul] }, - { rw [coe_mul_bot_of_neg hx, neg_top, ← coe_neg, coe_mul_bot_of_pos (neg_pos_of_neg hx)], }, }, - { norm_cast, exact neg_mul _ _, }, - { rcases lt_trichotomy 0 x with hx | rfl | hx, - { rw [coe_mul_top_of_pos hx, neg_top, ← coe_neg, coe_mul_top_of_neg (neg_neg_of_pos hx)], }, - { rw [coe_zero, zero_mul, neg_zero, zero_mul] }, - { rw [coe_mul_top_of_neg hx, neg_bot, ← coe_neg, coe_mul_top_of_pos (neg_pos_of_neg hx)], }, }, - { refl }, - { rcases lt_trichotomy 0 y with hy | rfl | hy, - { rw [top_mul_coe_of_pos hy, neg_top, bot_mul_coe_of_pos hy] }, - { rw [coe_zero, mul_zero, mul_zero, neg_zero] }, - { rw [top_mul_coe_of_neg hy, neg_top, neg_bot, bot_mul_coe_of_neg hy] } }, - { refl } + -- TODO: replace with `induction using` in Lean 4, which supports multiple premises + with_cases + { apply @induction₂ (λ x y, -x * y = -(x * y)) }; + propagate_tags { try { dsimp only} }, + case [top_top, bot_top, top_bot, bot_bot] { all_goals { refl } }, + case [top_zero, bot_zero, zero_top, zero_bot] + { all_goals { simp only [zero_mul, mul_zero, neg_zero] } }, + case coe_coe : x y { norm_cast, exact neg_mul _ _, }, + case pos_bot : x hx + { rw [coe_mul_bot_of_pos hx, neg_bot, ← coe_neg, coe_mul_bot_of_neg (neg_neg_of_pos hx)] }, + case neg_bot : x hx + { rw [coe_mul_bot_of_neg hx, neg_top, ← coe_neg, coe_mul_bot_of_pos (neg_pos_of_neg hx)] }, + case pos_top : x hx + { rw [coe_mul_top_of_pos hx, neg_top, ← coe_neg, coe_mul_top_of_neg (neg_neg_of_pos hx)] }, + case neg_top : x hx + { rw [coe_mul_top_of_neg hx, neg_bot, ← coe_neg, coe_mul_top_of_pos (neg_pos_of_neg hx)] }, + case top_pos : y hy { rw [top_mul_coe_of_pos hy, neg_top, bot_mul_coe_of_pos hy] }, + case top_neg : y hy { rw [top_mul_coe_of_neg hy, neg_top, neg_bot, bot_mul_coe_of_neg hy] }, + case bot_pos : y hy { rw [bot_mul_coe_of_pos hy, neg_bot, top_mul_coe_of_pos hy] }, + case bot_neg : y hy { rw [bot_mul_coe_of_neg hy, neg_bot, neg_top, top_mul_coe_of_neg hy] }, end instance : has_distrib_neg ereal := @@ -783,75 +826,62 @@ by rcases lt_trichotomy 0 x with hx | rfl | hx; simp [abs_def] @[simp] lemma abs_mul (x y : ereal) : (x * y).abs = x.abs * y.abs := begin - symmetry, - induction x using ereal.rec; induction y using ereal.rec, - { refl }, - { rcases lt_trichotomy 0 y with hy|rfl|hy, - { simp only [bot_mul_coe_of_pos hy, hy.ne', abs_bot, with_top.top_mul, ne.def, abs_eq_zero_iff, - coe_eq_zero, not_false_iff] }, - { simp only [coe_zero, abs_zero, mul_zero] }, - { simp only [bot_mul_coe_of_neg hy, hy.ne, abs_bot, with_top.top_mul, ne.def, abs_eq_zero_iff, - coe_eq_zero, not_false_iff, abs_top] } }, - { refl }, - { rcases lt_trichotomy 0 x with hx|rfl|hx, - { simp only [coe_mul_bot_of_pos hx, hx.ne', abs_bot, with_top.mul_top, ne.def, abs_eq_zero_iff, - coe_eq_zero, not_false_iff] }, - { simp only [coe_zero, abs_zero, zero_mul] }, - { simp only [coe_mul_bot_of_neg hx, hx.ne, abs_bot, with_top.mul_top, ne.def, abs_eq_zero_iff, - coe_eq_zero, not_false_iff, abs_top] } }, - { simp only [← coe_mul, ereal.abs, abs_mul, ennreal.of_real_mul (abs_nonneg _)] }, - { rcases lt_trichotomy 0 x with hx|rfl|hx, - { simp only [coe_mul_top_of_pos hx, hx.ne', with_top.mul_top, ne.def, abs_eq_zero_iff, - coe_eq_zero, not_false_iff, abs_top] }, - { simp only [coe_zero, abs_zero, zero_mul] }, - { simp only [coe_mul_top_of_neg hx, hx.ne, abs_bot, with_top.mul_top, ne.def, abs_eq_zero_iff, - coe_eq_zero, not_false_iff, abs_top] } }, - { refl }, - { rcases lt_trichotomy 0 y with hy|rfl|hy, - { simp only [top_mul_coe_of_pos hy, hy.ne', with_top.top_mul, ne.def, abs_eq_zero_iff, - coe_eq_zero, not_false_iff, abs_top] }, - { simp only [coe_zero, abs_zero, mul_zero] }, - { simp only [top_mul_coe_of_neg hy, hy.ne, abs_bot, with_top.top_mul, ne.def, abs_eq_zero_iff, - coe_eq_zero, not_false_iff, abs_top] } }, - { refl } + -- TODO: replace with `induction using` in Lean 4, which supports multiple premises + with_cases + { apply @induction₂ (λ x y, (x * y).abs = x.abs * y.abs) }; + propagate_tags { try { dsimp only} }, + case [top_top, bot_top, top_bot, bot_bot] { all_goals { refl } }, + case [top_zero, bot_zero, zero_top, zero_bot] { all_goals { simp only [zero_mul, mul_zero, + abs_zero] } }, + case coe_coe : x y { simp only [← coe_mul, ereal.abs, abs_mul, + ennreal.of_real_mul (abs_nonneg _)], }, + case pos_bot : x hx { simp only [coe_mul_bot_of_pos hx, hx.ne', abs_bot, with_top.mul_top, ne.def, + abs_eq_zero_iff, coe_eq_zero, not_false_iff] }, + case neg_bot : x hx { simp only [coe_mul_bot_of_neg hx, hx.ne, abs_bot, with_top.mul_top, ne.def, + abs_eq_zero_iff, coe_eq_zero, not_false_iff, abs_top] }, + case pos_top : x hx { simp only [coe_mul_top_of_pos hx, hx.ne', with_top.mul_top, ne.def, + abs_eq_zero_iff, coe_eq_zero, not_false_iff, abs_top] }, + case neg_top : x hx { simp only [coe_mul_top_of_neg hx, hx.ne, abs_bot, with_top.mul_top, ne.def, + abs_eq_zero_iff, coe_eq_zero, not_false_iff, abs_top] }, + case top_pos : y hy { simp only [top_mul_coe_of_pos hy, hy.ne', with_top.top_mul, ne.def, + abs_eq_zero_iff, coe_eq_zero, not_false_iff, abs_top] }, + case top_neg : y hy { simp only [top_mul_coe_of_neg hy, hy.ne, abs_bot, with_top.top_mul, ne.def, + abs_eq_zero_iff, coe_eq_zero, not_false_iff, abs_top] }, + case bot_pos : y hy { simp only [bot_mul_coe_of_pos hy, hy.ne', abs_bot, with_top.top_mul, ne.def, + abs_eq_zero_iff, coe_eq_zero, not_false_iff] }, + case bot_neg : y hy { simp only [bot_mul_coe_of_neg hy, hy.ne, abs_bot, with_top.top_mul, ne.def, + abs_eq_zero_iff, coe_eq_zero, not_false_iff, abs_top] }, end /-! ### Sign -/ -@[simp] lemma sign_top : sign (⊤ : ereal) = sign_type.pos := rfl -@[simp] lemma sign_bot : sign (⊥ : ereal) = sign_type.neg := rfl +@[simp] lemma sign_top : sign (⊤ : ereal) = 1 := rfl +@[simp] lemma sign_bot : sign (⊥ : ereal) = -1 := rfl @[simp] lemma sign_coe (x : ℝ) : sign (x : ereal) = sign x := by simp only [sign, order_hom.coe_fun_mk, ereal.coe_pos, ereal.coe_neg'] @[simp] lemma sign_mul (x y : ereal) : sign (x * y) = sign x * sign y := begin - induction x using ereal.rec; induction y using ereal.rec, - { refl }, - { rcases lt_trichotomy 0 y with hy|rfl|hy, - { simp only [bot_mul_coe_of_pos hy, hy, sign_coe, sign_pos, mul_one] }, - { simp only [coe_zero, mul_zero, sign_zero] }, - { simp only [bot_mul_coe_of_neg hy, hy, sign_top, sign_type.pos_eq_one, sign_bot, - sign_type.neg_eq_neg_one, sign_coe, sign_neg, mul_neg, mul_one, neg_neg] } }, - { refl }, - { rcases lt_trichotomy 0 x with hx|rfl|hx, - { simp only [coe_mul_bot_of_pos hx, hx, sign_bot, sign_type.neg_eq_neg_one, sign_coe, sign_pos, - mul_neg, mul_one] }, - { simp only [coe_zero, zero_mul, sign_zero] }, - { simp only [coe_mul_bot_of_neg hx, hx, sign_top, sign_type.pos_eq_one, sign_coe, sign_neg, - sign_bot, sign_type.neg_eq_neg_one, mul_neg, mul_one, neg_neg] } }, - { simp only [← coe_mul, sign_coe, sign_mul] }, - { rcases lt_trichotomy 0 x with hx|rfl|hx, - { simp only [coe_mul_top_of_pos hx, hx, sign_coe, sign_pos, mul_one, zero_lt_top]}, - { simp only [coe_zero, zero_mul, sign_zero] }, - { simp only [coe_mul_top_of_neg hx, hx, sign_top, sign_type.pos_eq_one, sign_coe, sign_neg, - sign_bot, sign_type.neg_eq_neg_one, mul_one] } }, - { refl }, - { rcases lt_trichotomy 0 y with hy|rfl|hy, - { simp only [top_mul_coe_of_pos hy, hy, sign_coe, sign_pos, mul_one] }, - { simp only [coe_zero, mul_zero, sign_zero] }, - { simp only [top_mul_coe_of_neg hy, hy, sign_top, sign_type.pos_eq_one, sign_bot, - sign_type.neg_eq_neg_one, sign_coe, sign_neg, mul_neg, mul_one]} }, - { refl } + -- TODO: replace with `induction using` in Lean 4, which supports multiple premises + with_cases + { apply @induction₂ (λ x y, sign (x * y) = sign x * sign y) }; + propagate_tags { try { dsimp only} }, + case [top_top, bot_top, top_bot, bot_bot] { all_goals { refl } }, + case [top_zero, bot_zero, zero_top, zero_bot] { all_goals { simp only [zero_mul, mul_zero, + sign_zero] } }, + case coe_coe : x y { simp only [← coe_mul, sign_coe, sign_mul], }, + case pos_bot : x hx { simp_rw [coe_mul_bot_of_pos hx, sign_coe, sign_pos hx, one_mul] }, + case neg_bot : x hx { simp_rw [coe_mul_bot_of_neg hx, sign_coe, sign_neg hx, sign_top, sign_bot, + neg_one_mul, neg_neg] }, + case pos_top : x hx { simp_rw [coe_mul_top_of_pos hx, sign_coe, sign_pos hx, one_mul] }, + case neg_top : x hx { simp_rw [coe_mul_top_of_neg hx, sign_coe, sign_neg hx, sign_top, sign_bot, + mul_one] }, + case top_pos : y hy { simp_rw [top_mul_coe_of_pos hy, sign_coe, sign_pos hy, mul_one] }, + case top_neg : y hy { simp_rw [top_mul_coe_of_neg hy, sign_coe, sign_neg hy, sign_top, sign_bot, + one_mul] }, + case bot_pos : y hy { simp_rw [bot_mul_coe_of_pos hy, sign_coe, sign_pos hy, mul_one] }, + case bot_neg : y hy { simp_rw [bot_mul_coe_of_neg hy, sign_coe, sign_neg hy, sign_top, sign_bot, + neg_one_mul, neg_neg] }, end lemma sign_mul_abs (x : ereal) : From d90149e913bf65828b011379540c3379e01105fd Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 24 Feb 2023 11:52:57 +0000 Subject: [PATCH 0520/1291] feat(algebra/quaternion): add `quaternion.im` for the skew-adjoint part (#18456) As requested by @YaelDillies, as prework for #18349 where I need to split the quaternion into real and imaginary parts. Only one marginally interesting lemma is included, `im_sq` which says `q.im^2 = -norm_sq q.im` --- src/algebra/quaternion.lean | 62 +++++++++++++++++++++++++++++++++--- src/analysis/quaternion.lean | 3 ++ 2 files changed, 61 insertions(+), 4 deletions(-) diff --git a/src/algebra/quaternion.lean b/src/algebra/quaternion.lean index 8ecb6bf2fb2c8..9939c2410e6b1 100644 --- a/src/algebra/quaternion.lean +++ b/src/algebra/quaternion.lean @@ -84,12 +84,21 @@ def equiv_tuple {R : Type*} (c₁ c₂ : R) : ℍ[R, c₁, c₂] ≃ (fin 4 → variables {S T R : Type*} [comm_ring R] {c₁ c₂ : R} (r x y z : R) (a b c : ℍ[R, c₁, c₂]) +/-- The imaginary part of a quaternion. -/ +def im (x : ℍ[R, c₁, c₂]) : ℍ[R, c₁, c₂] := ⟨0, x.im_i, x.im_j, x.im_k⟩ + +@[simp] lemma im_re : a.im.re = 0 := rfl +@[simp] lemma im_im_i : a.im.im_i = a.im_i := rfl +@[simp] lemma im_im_j : a.im.im_j = a.im_j := rfl +@[simp] lemma im_im_k : a.im.im_k = a.im_k := rfl +@[simp] lemma im_idem : a.im.im = a.im := rfl + instance : has_coe_t R (ℍ[R, c₁, c₂]) := ⟨λ x, ⟨x, 0, 0, 0⟩⟩ -@[simp] lemma coe_re : (x : ℍ[R, c₁, c₂]).re = x := rfl -@[simp] lemma coe_im_i : (x : ℍ[R, c₁, c₂]).im_i = 0 := rfl -@[simp] lemma coe_im_j : (x : ℍ[R, c₁, c₂]).im_j = 0 := rfl -@[simp] lemma coe_im_k : (x : ℍ[R, c₁, c₂]).im_k = 0 := rfl +@[simp, norm_cast] lemma coe_re : (x : ℍ[R, c₁, c₂]).re = x := rfl +@[simp, norm_cast] lemma coe_im_i : (x : ℍ[R, c₁, c₂]).im_i = 0 := rfl +@[simp, norm_cast] lemma coe_im_j : (x : ℍ[R, c₁, c₂]).im_j = 0 := rfl +@[simp, norm_cast] lemma coe_im_k : (x : ℍ[R, c₁, c₂]).im_k = 0 := rfl lemma coe_injective : function.injective (coe : R → ℍ[R, c₁, c₂]) := λ x y h, congr_arg re h @@ -131,6 +140,17 @@ by ext; simp (mk a₁ a₂ a₃ a₄ : ℍ[R, c₁, c₂]) - mk b₁ b₂ b₃ b₄ = mk (a₁ - b₁) (a₂ - b₂) (a₃ - b₃) (a₄ - b₄) := rfl +@[simp, norm_cast] lemma coe_im : (x : ℍ[R, c₁, c₂]).im = 0 := rfl + +@[simp] lemma re_add_im : ↑a.re + a.im = a := +ext _ _ (add_zero _) (zero_add _) (zero_add _) (zero_add _) + +@[simp] lemma sub_self_im : a - a.im = a.re := +ext _ _ (sub_zero _) (sub_self _) (sub_self _) (sub_self _) + +@[simp] lemma sub_self_re : a - a.re = a.im := +ext _ _ (sub_self _) (sub_zero _) (sub_zero _) (sub_zero _) + /-- Multiplication is given by * `1 * x = x * 1 = x`; @@ -210,12 +230,14 @@ instance : add_group_with_one ℍ[R, c₁, c₂] := @[simp, norm_cast] lemma nat_cast_im_i (n : ℕ) : (n : ℍ[R, c₁, c₂]).im_i = 0 := rfl @[simp, norm_cast] lemma nat_cast_im_j (n : ℕ) : (n : ℍ[R, c₁, c₂]).im_j = 0 := rfl @[simp, norm_cast] lemma nat_cast_im_k (n : ℕ) : (n : ℍ[R, c₁, c₂]).im_k = 0 := rfl +@[simp, norm_cast] lemma nat_cast_im (n : ℕ) : (n : ℍ[R, c₁, c₂]).im = 0 := rfl @[norm_cast] lemma coe_nat_cast (n : ℕ) : ↑(n : R) = (n : ℍ[R, c₁, c₂]) := rfl @[simp, norm_cast] lemma int_cast_re (z : ℤ) : (z : ℍ[R, c₁, c₂]).re = z := rfl @[simp, norm_cast] lemma int_cast_im_i (z : ℤ) : (z : ℍ[R, c₁, c₂]).im_i = 0 := rfl @[simp, norm_cast] lemma int_cast_im_j (z : ℤ) : (z : ℍ[R, c₁, c₂]).im_j = 0 := rfl @[simp, norm_cast] lemma int_cast_im_k (z : ℤ) : (z : ℍ[R, c₁, c₂]).im_k = 0 := rfl +@[simp, norm_cast] lemma int_cast_im (z : ℤ) : (z : ℍ[R, c₁, c₂]).im = 0 := rfl @[norm_cast] lemma coe_int_cast (z : ℤ) : ↑(z : R) = (z : ℍ[R, c₁, c₂]) := rfl instance : ring ℍ[R, c₁, c₂] := @@ -328,6 +350,7 @@ linear_equiv.of_involutive @[simp] lemma im_i_conj : (conj a).im_i = - a.im_i := rfl @[simp] lemma im_j_conj : (conj a).im_j = - a.im_j := rfl @[simp] lemma im_k_conj : (conj a).im_k = - a.im_k := rfl +@[simp] lemma im_conj : (conj a).im = - a.im := ext _ _ neg_zero.symm rfl rfl rfl @[simp] lemma conj_mk (a₁ a₂ a₃ a₄ : R) : conj (mk a₁ a₂ a₃ a₄ : ℍ[R, c₁, c₂]) = ⟨a₁, -a₂, -a₃, -a₄⟩ := @@ -372,6 +395,8 @@ calc a.conj * b.conj = (b * a).conj : (conj_mul b a).symm @[simp, norm_cast] lemma conj_coe : conj (x : ℍ[R, c₁, c₂]) = x := by ext; simp +@[simp] lemma conj_im : conj a.im = - a.im := im_conj _ + @[simp, norm_cast] lemma conj_nat_cast (n : ℕ) : conj (n : ℍ[R, c₁, c₂]) = n := by rw [←coe_nat_cast, conj_coe] @[simp, norm_cast] lemma conj_int_cast (z : ℤ) : conj (z : ℍ[R, c₁, c₂]) = z := @@ -476,39 +501,58 @@ lemma ext_iff {a b : ℍ[R]} : a = b ↔ a.re = b.re ∧ a.im_i = b.im_i ∧ a.im_j = b.im_j ∧ a.im_k = b.im_k := quaternion_algebra.ext_iff a b +/-- The imaginary part of a quaternion. -/ +def im (x : ℍ[R]) : ℍ[R] := x.im + +@[simp] lemma im_re : a.im.re = 0 := rfl +@[simp] lemma im_im_i : a.im.im_i = a.im_i := rfl +@[simp] lemma im_im_j : a.im.im_j = a.im_j := rfl +@[simp] lemma im_im_k : a.im.im_k = a.im_k := rfl +@[simp] lemma im_idem : a.im.im = a.im := rfl + +@[simp] lemma re_add_im : ↑a.re + a.im = a := a.re_add_im +@[simp] lemma sub_self_im : a - a.im = a.re := a.sub_self_im +@[simp] lemma sub_self_re : a - a.re = a.im := a.sub_self_re + @[simp, norm_cast] lemma coe_re : (x : ℍ[R]).re = x := rfl @[simp, norm_cast] lemma coe_im_i : (x : ℍ[R]).im_i = 0 := rfl @[simp, norm_cast] lemma coe_im_j : (x : ℍ[R]).im_j = 0 := rfl @[simp, norm_cast] lemma coe_im_k : (x : ℍ[R]).im_k = 0 := rfl +@[simp, norm_cast] lemma coe_im : (x : ℍ[R]).im = 0 := rfl @[simp] lemma zero_re : (0 : ℍ[R]).re = 0 := rfl @[simp] lemma zero_im_i : (0 : ℍ[R]).im_i = 0 := rfl @[simp] lemma zero_im_j : (0 : ℍ[R]).im_j = 0 := rfl @[simp] lemma zero_im_k : (0 : ℍ[R]).im_k = 0 := rfl +@[simp] lemma zero_im : (0 : ℍ[R]).im = 0 := rfl @[simp, norm_cast] lemma coe_zero : ((0 : R) : ℍ[R]) = 0 := rfl @[simp] lemma one_re : (1 : ℍ[R]).re = 1 := rfl @[simp] lemma one_im_i : (1 : ℍ[R]).im_i = 0 := rfl @[simp] lemma one_im_j : (1 : ℍ[R]).im_j = 0 := rfl @[simp] lemma one_im_k : (1 : ℍ[R]).im_k = 0 := rfl +@[simp] lemma one_im : (1 : ℍ[R]).im = 0 := rfl @[simp, norm_cast] lemma coe_one : ((1 : R) : ℍ[R]) = 1 := rfl @[simp] lemma add_re : (a + b).re = a.re + b.re := rfl @[simp] lemma add_im_i : (a + b).im_i = a.im_i + b.im_i := rfl @[simp] lemma add_im_j : (a + b).im_j = a.im_j + b.im_j := rfl @[simp] lemma add_im_k : (a + b).im_k = a.im_k + b.im_k := rfl +@[simp] lemma add_im : (a + b).im = a.im + b.im := ext _ _ (add_zero _).symm rfl rfl rfl @[simp, norm_cast] lemma coe_add : ((x + y : R) : ℍ[R]) = x + y := quaternion_algebra.coe_add x y @[simp] lemma neg_re : (-a).re = -a.re := rfl @[simp] lemma neg_im_i : (-a).im_i = -a.im_i := rfl @[simp] lemma neg_im_j : (-a).im_j = -a.im_j := rfl @[simp] lemma neg_im_k : (-a).im_k = -a.im_k := rfl +@[simp] lemma neg_im : (-a).im = -a.im := ext _ _ neg_zero.symm rfl rfl rfl @[simp, norm_cast] lemma coe_neg : ((-x : R) : ℍ[R]) = -x := quaternion_algebra.coe_neg x @[simp] lemma sub_re : (a - b).re = a.re - b.re := rfl @[simp] lemma sub_im_i : (a - b).im_i = a.im_i - b.im_i := rfl @[simp] lemma sub_im_j : (a - b).im_j = a.im_j - b.im_j := rfl @[simp] lemma sub_im_k : (a - b).im_k = a.im_k - b.im_k := rfl +@[simp] lemma sub_im : (a - b).im = a.im - b.im := ext _ _ (sub_zero _).symm rfl rfl rfl @[simp, norm_cast] lemma coe_sub : ((x - y : R) : ℍ[R]) = x - y := quaternion_algebra.coe_sub x y @[simp] lemma mul_re : @@ -540,12 +584,14 @@ quaternion_algebra.coe_pow x n @[simp, norm_cast] lemma nat_cast_im_i (n : ℕ) : (n : ℍ[R]).im_i = 0 := rfl @[simp, norm_cast] lemma nat_cast_im_j (n : ℕ) : (n : ℍ[R]).im_j = 0 := rfl @[simp, norm_cast] lemma nat_cast_im_k (n : ℕ) : (n : ℍ[R]).im_k = 0 := rfl +@[simp, norm_cast] lemma nat_cast_im (n : ℕ) : (n : ℍ[R]).im = 0 := rfl @[norm_cast] lemma coe_nat_cast (n : ℕ) : ↑(n : R) = (n : ℍ[R]) := rfl @[simp, norm_cast] lemma int_cast_re (z : ℤ) : (z : ℍ[R]).re = z := rfl @[simp, norm_cast] lemma int_cast_im_i (z : ℤ) : (z : ℍ[R]).im_i = 0 := rfl @[simp, norm_cast] lemma int_cast_im_j (z : ℤ) : (z : ℍ[R]).im_j = 0 := rfl @[simp, norm_cast] lemma int_cast_im_k (z : ℤ) : (z : ℍ[R]).im_k = 0 := rfl +@[simp, norm_cast] lemma int_cast_im (z : ℤ) : (z : ℍ[R]).im = 0 := rfl @[norm_cast] lemma coe_int_cast (z : ℤ) : ↑(z : R) = (z : ℍ[R]) := rfl lemma coe_injective : function.injective (coe : R → ℍ[R]) := quaternion_algebra.coe_injective @@ -556,6 +602,8 @@ lemma coe_injective : function.injective (coe : R → ℍ[R]) := quaternion_alge @[simp] lemma smul_im_i [has_smul S R] (s : S) : (s • a).im_i = s • a.im_i := rfl @[simp] lemma smul_im_j [has_smul S R] (s : S) : (s • a).im_j = s • a.im_j := rfl @[simp] lemma smul_im_k [has_smul S R] (s : S) : (s • a).im_k = s • a.im_k := rfl +@[simp] lemma smul_im [smul_zero_class S R] (s : S) : (s • a).im = s • a.im := +ext _ _ (smul_zero _).symm rfl rfl rfl @[simp, norm_cast] lemma coe_smul [smul_zero_class S R] (s : S) (r : R) : (↑(s • r) : ℍ[R]) = s • ↑r := @@ -589,6 +637,7 @@ def conj : ℍ[R] ≃ₗ[R] ℍ[R] := quaternion_algebra.conj @[simp] lemma conj_im_i : a.conj.im_i = - a.im_i := rfl @[simp] lemma conj_im_j : a.conj.im_j = - a.im_j := rfl @[simp] lemma conj_im_k : a.conj.im_k = - a.im_k := rfl +@[simp] lemma conj_im : a.conj.im = - a.im := a.im_conj @[simp] lemma conj_conj : a.conj.conj = a := a.conj_conj @@ -620,6 +669,7 @@ quaternion_algebra.commute_conj_conj h alias commute_conj_conj ← commute.quaternion_conj @[simp, norm_cast] lemma conj_coe : conj (x : ℍ[R]) = x := quaternion_algebra.conj_coe x +@[simp] lemma im_conj : a.im.conj = - a.im := quaternion_algebra.im_conj _ @[simp, norm_cast] lemma conj_nat_cast (n : ℕ) : conj (n : ℍ[R]) = n := quaternion_algebra.conj_nat_cast _ @@ -692,6 +742,9 @@ lemma self_mul_conj : a * a.conj = norm_sq a := by rw [mul_conj_eq_coe, norm_sq_ lemma conj_mul_self : a.conj * a = norm_sq a := by rw [← a.commute_self_conj.eq, self_mul_conj] +lemma im_sq : a.im^2 = -norm_sq a.im := +by simp_rw [sq, ←conj_mul_self, im_conj, neg_mul, neg_neg] + lemma coe_norm_sq_add : (norm_sq (a + b) : ℍ[R]) = norm_sq a + a * b.conj + b * a.conj + norm_sq b := by simp [← self_mul_conj, mul_add, add_mul, add_assoc] @@ -774,6 +827,7 @@ instance : division_ring ℍ[R] := @[simp, norm_cast] lemma rat_cast_im_i (q : ℚ) : (q : ℍ[R]).im_i = 0 := rfl @[simp, norm_cast] lemma rat_cast_im_j (q : ℚ) : (q : ℍ[R]).im_j = 0 := rfl @[simp, norm_cast] lemma rat_cast_im_k (q : ℚ) : (q : ℍ[R]).im_k = 0 := rfl +@[simp, norm_cast] lemma rat_cast_im (q : ℚ) : (q : ℍ[R]).im = 0 := rfl @[norm_cast] lemma coe_rat_cast (q : ℚ) : ↑(q : R) = (q : ℍ[R]) := rfl lemma conj_inv : conj (a⁻¹) = (conj a)⁻¹ := star_inv' a diff --git a/src/analysis/quaternion.lean b/src/analysis/quaternion.lean index d1c6d0ed97037..3b8116325036b 100644 --- a/src/analysis/quaternion.lean +++ b/src/analysis/quaternion.lean @@ -147,6 +147,9 @@ by simpa [←norm_sq_eq_norm_sq] @[continuity] lemma continuous_im_k : continuous (λ q : ℍ, q.im_k) := (continuous_apply 3).comp linear_isometry_equiv_tuple.continuous +@[continuity] lemma continuous_im : continuous (λ q : ℍ, q.im) := +by simpa only [←sub_self_re] using continuous_id.sub (continuous_coe.comp continuous_re) + instance : complete_space ℍ := begin have : uniform_embedding linear_isometry_equiv_tuple.to_linear_equiv.to_equiv.symm := From b2ff9a3d7a15fd5b0f060b135421d6a89a999c2f Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Fri, 24 Feb 2023 11:52:58 +0000 Subject: [PATCH 0521/1291] chore(data/real/nnreal): drop some lemmas (#18492) These lemmas are now available for any semifield, so we don't need them. All these lemmas are `@[deprecated]` in the Mathlib 4 port and will be removed in the forward-port of this PR. --- src/analysis/analytic/composition.lean | 2 +- src/analysis/calculus/inverse.lean | 6 ++-- src/analysis/normed_space/star/mul.lean | 2 +- src/analysis/specific_limits/basic.lean | 2 +- src/data/real/nnreal.lean | 32 +------------------ .../covering/differentiation.lean | 2 +- .../covering/liminf_limsup.lean | 2 +- .../decomposition/lebesgue.lean | 2 +- .../conditional_expectation/real.lean | 2 +- src/measure_theory/function/jacobian.lean | 2 +- .../integral/riesz_markov_kakutani.lean | 6 ++-- .../metric_space/hausdorff_distance.lean | 2 +- .../metric_space/metrizable_uniformity.lean | 2 +- 13 files changed, 17 insertions(+), 47 deletions(-) diff --git a/src/analysis/analytic/composition.lean b/src/analysis/analytic/composition.lean index 87bc85aa02f33..a0e417d0692ad 100644 --- a/src/analysis/analytic/composition.lean +++ b/src/analysis/analytic/composition.lean @@ -453,7 +453,7 @@ begin { 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)⁻¹, - have r0_pos : 0 < r0 := nnreal.inv_pos.2 (mul_pos zero_lt_four (zero_lt_one.trans_le hCp1)), + have r0_pos : 0 < r0 := inv_pos.2 (mul_pos zero_lt_four (zero_lt_one.trans_le hCp1)), 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), diff --git a/src/analysis/calculus/inverse.lean b/src/analysis/calculus/inverse.lean index fa25e0d52e932..6f928e563362d 100644 --- a/src/analysis/calculus/inverse.lean +++ b/src/analysis/calculus/inverse.lean @@ -561,7 +561,7 @@ begin let f'symm := f'.nonlinear_right_inverse_of_surjective h, set c : ℝ≥0 := f'symm.nnnorm⁻¹ / 2 with hc, have f'symm_pos : 0 < f'symm.nnnorm := f'.nonlinear_right_inverse_of_surjective_nnnorm_pos h, - have cpos : 0 < c, by simp [hc, nnreal.half_pos, nnreal.inv_pos, f'symm_pos], + have cpos : 0 < c, by simp [hc, half_pos, inv_pos, f'symm_pos], obtain ⟨s, s_nhds, hs⟩ : ∃ s ∈ 𝓝 a, approximates_linear_on f f' s c := hf.approximates_deriv_on_nhds (or.inr cpos), apply hs.map_nhds_eq f'symm s_nhds (or.inr (nnreal.half_lt_self _)), @@ -577,7 +577,7 @@ begin refine ((nhds_basis_opens a).exists_iff _).1 _, exact (λ s t, approximates_linear_on.mono_set), exact (hf.approximates_deriv_on_nhds $ f'.subsingleton_or_nnnorm_symm_pos.imp id $ - λ hf', nnreal.half_pos $ nnreal.inv_pos.2 $ hf') + λ hf', half_pos $ inv_pos.2 hf') end include cs @@ -593,7 +593,7 @@ approximates_linear_on.to_local_homeomorph f (classical.some hf.approximates_deriv_on_open_nhds) (classical.some_spec hf.approximates_deriv_on_open_nhds).snd (f'.subsingleton_or_nnnorm_symm_pos.imp id $ λ hf', nnreal.half_lt_self $ ne_of_gt $ - nnreal.inv_pos.2 $ hf') + inv_pos.2 hf') (classical.some_spec hf.approximates_deriv_on_open_nhds).fst.2 variable {f} diff --git a/src/analysis/normed_space/star/mul.lean b/src/analysis/normed_space/star/mul.lean index 55dbf1ba90df0..46e1f84950911 100644 --- a/src/analysis/normed_space/star/mul.lean +++ b/src/analysis/normed_space/star/mul.lean @@ -29,7 +29,7 @@ begin { have ha : 0 < ‖a‖₊ := zero_le'.trans_lt hr, rw [←inv_inv (‖a‖₊), nnreal.lt_inv_iff_mul_lt (inv_ne_zero ha.ne')] at hr, obtain ⟨k, hk₁, hk₂⟩ := normed_field.exists_lt_nnnorm_lt 𝕜 (mul_lt_mul_of_pos_right hr $ - nnreal.inv_pos.2 ha), + inv_pos.2 ha), refine ⟨_, ⟨k • star a, _, rfl⟩, _⟩, { simpa only [mem_closed_ball_zero_iff, norm_smul, one_mul, norm_star] using (nnreal.le_inv_iff_mul_le ha.ne').1 (one_mul ‖a‖₊⁻¹ ▸ hk₂.le : ‖k‖₊ ≤ ‖a‖₊⁻¹) }, diff --git a/src/analysis/specific_limits/basic.lean b/src/analysis/specific_limits/basic.lean index 6213b01d8c5d4..28f25521d4b77 100644 --- a/src/analysis/specific_limits/basic.lean +++ b/src/analysis/specific_limits/basic.lean @@ -503,7 +503,7 @@ begin lift w to ι → ℝ≥0 using hw, rcases exists_pos_sum_of_countable hε ι with ⟨δ', Hpos, Hsum⟩, have : ∀ i, 0 < max 1 (w i), from λ i, zero_lt_one.trans_le (le_max_left _ _), - refine ⟨λ i, δ' i / max 1 (w i), λ i, nnreal.div_pos (Hpos _) (this i), _⟩, + refine ⟨λ i, δ' i / max 1 (w i), λ i, div_pos (Hpos _) (this i), _⟩, refine lt_of_le_of_lt (ennreal.tsum_le_tsum $ λ i, _) Hsum, rw [coe_div (this i).ne'], refine mul_le_of_le_div' (ennreal.mul_le_mul le_rfl $ ennreal.inv_le_inv.2 _), diff --git a/src/data/real/nnreal.lean b/src/data/real/nnreal.lean index 43084ff20997a..09a78d5f3819c 100644 --- a/src/data/real/nnreal.lean +++ b/src/data/real/nnreal.lean @@ -579,12 +579,6 @@ lemma sum_div {ι} (s : finset ι) (f : ι → ℝ≥0) (b : ℝ≥0) : (∑ i in s, f i) / b = ∑ i in s, (f i / b) := finset.sum_div -@[simp] lemma inv_pos {r : ℝ≥0} : 0 < r⁻¹ ↔ 0 < r := inv_pos - -lemma div_pos {r p : ℝ≥0} (hr : 0 < r) (hp : 0 < p) : 0 < r / p := div_pos hr hp - -lemma div_self_le (r : ℝ≥0) : r / r ≤ 1 := div_self_le_one r - @[simp] lemma inv_le {r p : ℝ≥0} (h : r ≠ 0) : r⁻¹ ≤ p ↔ 1 ≤ r * p := by rw [← mul_le_mul_left (pos_iff_ne_zero.2 h), mul_inv_cancel h] @@ -661,36 +655,16 @@ le_of_forall_ge_of_dense $ assume a ha, have (a * x⁻¹) * x ≤ y, from h _ this, by rwa [mul_assoc, inv_mul_cancel hx, mul_one] at this -lemma div_add_div_same (a b c : ℝ≥0) : a / c + b / c = (a + b) / c := div_add_div_same _ _ _ - -lemma half_pos {a : ℝ≥0} (h : 0 < a) : 0 < a / 2 := half_pos h - -lemma add_halves (a : ℝ≥0) : a / 2 + a / 2 = a := add_halves _ - lemma half_le_self (a : ℝ≥0) : a / 2 ≤ a := half_le_self bot_le lemma half_lt_self {a : ℝ≥0} (h : a ≠ 0) : a / 2 < a := half_lt_self h.bot_lt -lemma two_inv_lt_one : (2⁻¹:ℝ≥0) < 1 := two_inv_lt_one - lemma div_lt_one_of_lt {a b : ℝ≥0} (h : a < b) : a / b < 1 := begin rwa [div_lt_iff, one_mul], exact ne_of_gt (lt_of_le_of_lt (zero_le _) h) end -@[field_simps] lemma div_add_div (a : ℝ≥0) {b : ℝ≥0} (c : ℝ≥0) {d : ℝ≥0} - (hb : b ≠ 0) (hd : d ≠ 0) : a / b + c / d = (a * d + b * c) / (b * d) := -div_add_div _ _ hb hd - -@[field_simps] lemma add_div' (a b c : ℝ≥0) (hc : c ≠ 0) : - b + a / c = (b * c + a) / c := -add_div' _ _ _ hc - -@[field_simps] lemma div_add' (a b c : ℝ≥0) (hc : c ≠ 0) : - a / c + b = (a + b * c) / c := -div_add' _ _ _ hc - lemma _root_.real.to_nnreal_inv {x : ℝ} : real.to_nnreal x⁻¹ = (real.to_nnreal x)⁻¹ := begin @@ -712,8 +686,6 @@ by rw [div_eq_inv_mul, div_eq_inv_mul, real.to_nnreal_mul (inv_nonneg.2 hy), rea lemma inv_lt_one_iff {x : ℝ≥0} (hx : x ≠ 0) : x⁻¹ < 1 ↔ 1 < x := by rwa [← one_div, div_lt_iff hx, one_mul] -lemma inv_lt_one {x : ℝ≥0} (hx : 1 < x) : x⁻¹ < 1 := inv_lt_one hx - lemma zpow_pos {x : ℝ≥0} (hx : x ≠ 0) (n : ℤ) : 0 < x ^ n := begin cases n, @@ -721,10 +693,8 @@ begin { simp [pow_pos hx.bot_lt _] } end -lemma inv_lt_inv_iff {x y : ℝ≥0} (hx : x ≠ 0) (hy : y ≠ 0) : y⁻¹ < x⁻¹ ↔ x < y := inv_lt_inv₀ hy hx - lemma inv_lt_inv {x y : ℝ≥0} (hx : x ≠ 0) (h : x < y) : y⁻¹ < x⁻¹ := -(inv_lt_inv_iff hx ((bot_le.trans_lt h).ne')).2 h +inv_lt_inv_of_lt hx.bot_lt h end inv diff --git a/src/measure_theory/covering/differentiation.lean b/src/measure_theory/covering/differentiation.lean index 7aa3813ee22a3..e70665357ccb9 100644 --- a/src/measure_theory/covering/differentiation.lean +++ b/src/measure_theory/covering/differentiation.lean @@ -594,7 +594,7 @@ begin conv_rhs { rw ← mul_one (t^ n) }, refine mul_lt_mul' le_rfl _ (zero_le _) (nnreal.zpow_pos t_ne_zero' _), rw zpow_neg_one, - exact nnreal.inv_lt_one ht, + exact inv_lt_one ht, end }, calc ν s = ν (s ∩ f⁻¹' {0}) + ν (s ∩ f⁻¹' {∞}) + ∑' (n : ℤ), ν (s ∩ f⁻¹' (Ico (t^n) (t^(n+1)))) : measure_eq_measure_preimage_add_measure_tsum_Ico_zpow ν f_meas hs ht diff --git a/src/measure_theory/covering/liminf_limsup.lean b/src/measure_theory/covering/liminf_limsup.lean index 911c1c887e9f7..a53637132c1a2 100644 --- a/src/measure_theory/covering/liminf_limsup.lean +++ b/src/measure_theory/covering/liminf_limsup.lean @@ -105,7 +105,7 @@ begin replace hη' : 1 ≤ η := by simpa only [ennreal.one_le_coe_iff] using le_of_tendsto (hd' w (λ j, r₁ (f j)) hr $ eventually_of_forall hw') hη', exact (lt_self_iff_false _).mp (lt_of_lt_of_le hη hη'), }, - refine ⟨1 - C⁻¹, tsub_lt_self zero_lt_one (nnreal.inv_pos.mpr hC), _⟩, + refine ⟨1 - C⁻¹, tsub_lt_self zero_lt_one (inv_pos.mpr hC), _⟩, replace hC : C ≠ 0 := ne_of_gt hC, let b : ℕ → set α := λ j, closed_ball (w j) (M * r₁ (f j)), let B : ℕ → set α := λ j, closed_ball (w j) (r₁ (f j)), diff --git a/src/measure_theory/decomposition/lebesgue.lean b/src/measure_theory/decomposition/lebesgue.lean index 85b42905182f4..eb1a19d20703b 100644 --- a/src/measure_theory/decomposition/lebesgue.lean +++ b/src/measure_theory/decomposition/lebesgue.lean @@ -429,7 +429,7 @@ begin by_cases hb : 0 < νA, { suffices : ∀ b, 0 < b → μA ≤ b, { by_contra, - have h' := this (μA / 2) (nnreal.half_pos (zero_lt_iff.2 h)), + have h' := this (μA / 2) (half_pos (zero_lt_iff.2 h)), rw ← @not_not (μA ≤ μA / 2) at h', exact h' (not_le.2 (nnreal.half_lt_self h)) }, intros c hc, diff --git a/src/measure_theory/function/conditional_expectation/real.lean b/src/measure_theory/function/conditional_expectation/real.lean index 4fa0ccd6f1fd9..1d5a8929a2e18 100644 --- a/src/measure_theory/function/conditional_expectation/real.lean +++ b/src/measure_theory/function/conditional_expectation/real.lean @@ -212,7 +212,7 @@ begin obtain ⟨δ, hδ, h⟩ := hg.snorm_indicator_le μ le_rfl ennreal.one_ne_top hε, set C : ℝ≥0 := ⟨δ, hδ.le⟩⁻¹ * (snorm g 1 μ).to_nnreal with hC, have hCpos : 0 < C := - mul_pos (nnreal.inv_pos.2 hδ) (ennreal.to_nnreal_pos hne hg.snorm_lt_top.ne), + mul_pos (inv_pos.2 hδ) (ennreal.to_nnreal_pos hne hg.snorm_lt_top.ne), have : ∀ n, μ {x : α | C ≤ ‖μ[g | ℱ n] x‖₊} ≤ ennreal.of_real δ, { intro n, have := mul_meas_ge_le_pow_snorm' μ one_ne_zero ennreal.one_ne_top diff --git a/src/measure_theory/function/jacobian.lean b/src/measure_theory/function/jacobian.lean index 1052d71d844a7..999c80fd00745 100644 --- a/src/measure_theory/function/jacobian.lean +++ b/src/measure_theory/function/jacobian.lean @@ -415,7 +415,7 @@ begin { simp only [h, true_or, eventually_const] }, simp only [h, false_or], apply Iio_mem_nhds, - simpa only [h, false_or, nnreal.inv_pos] using B.subsingleton_or_nnnorm_symm_pos }, + simpa only [h, false_or, inv_pos] using B.subsingleton_or_nnnorm_symm_pos }, have L2 : ∀ᶠ δ in 𝓝 (0 : ℝ≥0), ‖(B.symm : E →L[ℝ] E)‖₊ * (‖(B.symm : E →L[ℝ] E)‖₊⁻¹ - δ)⁻¹ * δ < δ₀, { have : tendsto (λ δ, ‖(B.symm : E →L[ℝ] E)‖₊ * (‖(B.symm : E →L[ℝ] E)‖₊⁻¹ - δ)⁻¹ * δ) diff --git a/src/measure_theory/integral/riesz_markov_kakutani.lean b/src/measure_theory/integral/riesz_markov_kakutani.lean index 24aeedef3703a..aeff61a56f4f4 100644 --- a/src/measure_theory/integral/riesz_markov_kakutani.lean +++ b/src/measure_theory/integral/riesz_markov_kakutani.lean @@ -93,9 +93,9 @@ begin intros ε εpos, --get test functions s.t. `λ(Ki) ≤ Λfi ≤ λ(Ki) + ε/2, i=1,2` obtain ⟨f1, f_test_function_K1⟩ := exists_lt_riesz_content_aux_add_pos Λ K1 - (nnreal.half_pos εpos), + (half_pos εpos), obtain ⟨f2, f_test_function_K2⟩ := exists_lt_riesz_content_aux_add_pos Λ K2 - (nnreal.half_pos εpos), + (half_pos εpos), --let `f := f1 + f2` test function for the content of `K` have f_test_function_union : (∀ x ∈ (K1 ⊔ K2), (1 : ℝ≥0) ≤ (f1 + f2) x), { rintros x (x_in_K1 | x_in_K2), @@ -106,7 +106,7 @@ begin rw map_add, --use that `Λfi` are lower bounds for `λ(Ki) + ε/2` apply lt_of_lt_of_le (add_lt_add f_test_function_K1.right f_test_function_K2.right) (le_of_eq _), - rw [add_assoc, add_comm (ε/2), add_assoc, nnreal.add_halves ε, add_assoc], + rw [add_assoc, add_comm (ε/2), add_assoc, add_halves ε, add_assoc], end end riesz_subadditive diff --git a/src/topology/metric_space/hausdorff_distance.lean b/src/topology/metric_space/hausdorff_distance.lean index d2c2a35abd69d..91616d78d4785 100644 --- a/src/topology/metric_space/hausdorff_distance.lean +++ b/src/topology/metric_space/hausdorff_distance.lean @@ -1166,7 +1166,7 @@ begin refine (h x hx y hy).not_le _, calc edist x y ≤ edist z x + edist z y : edist_triangle_left _ _ _ ... ≤ ↑(r / 2) + ↑(r / 2) : add_le_add hzx.le hzy.le - ... = r : by rw [← ennreal.coe_add, nnreal.add_halves] + ... = r : by rw [← ennreal.coe_add, add_halves] end lemma _root_.disjoint.exists_cthickenings (hst : disjoint s t) (hs : is_compact s) diff --git a/src/topology/metric_space/metrizable_uniformity.lean b/src/topology/metric_space/metrizable_uniformity.lean index 599bd8385e659..bed0c90bf10f3 100644 --- a/src/topology/metric_space/metrizable_uniformity.lean +++ b/src/topology/metric_space/metrizable_uniformity.lean @@ -199,7 +199,7 @@ begin { intros x y, dsimp only [d], simp only [@symmetric_rel.mk_mem_comm _ _ (hU_symm _) x y] }, have hr : (1 / 2 : ℝ≥0) ∈ Ioo (0 : ℝ≥0) 1, - from ⟨nnreal.half_pos one_pos, nnreal.half_lt_self one_ne_zero⟩, + from ⟨half_pos one_pos, nnreal.half_lt_self one_ne_zero⟩, letI I := pseudo_metric_space.of_prenndist d (λ x, hd₀.2 (setoid.refl _)) hd_symm, have hdist_le : ∀ x y, dist x y ≤ d x y, from pseudo_metric_space.dist_of_prenndist_le _ _ _, From a8e7ac804fc39df0340c64906075787e0c90fa60 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Mar=C3=ADa=20In=C3=A9s=20de=20Frutos-Fern=C3=A1ndez?= Date: Fri, 24 Feb 2023 13:53:44 +0000 Subject: [PATCH 0522/1291] feat(ring_theory/dedekind_domain/finite_adele_ring): add finite_adele_ring (#14249) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We define the finite adèle ring of a Dedekind domain and show that it is a commutative ring. Co-authored-by: Oliver Nash --- .../dedekind_domain/adic_valuation.lean | 92 +++++++ .../dedekind_domain/finite_adele_ring.lean | 228 ++++++++++++++++++ src/topology/algebra/valued_field.lean | 3 + 3 files changed, 323 insertions(+) create mode 100644 src/ring_theory/dedekind_domain/finite_adele_ring.lean diff --git a/src/ring_theory/dedekind_domain/adic_valuation.lean b/src/ring_theory/dedekind_domain/adic_valuation.lean index ea551209d81bb..2dfb9e9d972a8 100644 --- a/src/ring_theory/dedekind_domain/adic_valuation.lean +++ b/src/ring_theory/dedekind_domain/adic_valuation.lean @@ -332,4 +332,96 @@ instance adic_completion.has_lift_t : has_lift_t K (v.adic_completion K) := /-- The ring of integers of `adic_completion`. -/ def adic_completion_integers : valuation_subring (v.adic_completion K) := valued.v.valuation_subring +instance : inhabited (adic_completion_integers K v) := ⟨0⟩ + +variables (R K) + +lemma mem_adic_completion_integers {x : v.adic_completion K} : + x ∈ v.adic_completion_integers K ↔ (valued.v x : ℤₘ₀) ≤ 1 := +iff.rfl + +section algebra_instances + +@[priority 100] instance adic_valued.has_uniform_continuous_const_smul' : + @has_uniform_continuous_const_smul R K v.adic_valued.to_uniform_space _ := +@has_uniform_continuous_const_smul_of_continuous_const_smul R K _ _ _ + v.adic_valued.to_uniform_space _ _ + +instance adic_valued.has_uniform_continuous_const_smul : + @has_uniform_continuous_const_smul K K v.adic_valued.to_uniform_space _ := +@ring.has_uniform_continuous_const_smul K _ v.adic_valued.to_uniform_space _ _ + +instance adic_completion.algebra' : algebra R (v.adic_completion K) := +@uniform_space.completion.algebra K _ v.adic_valued.to_uniform_space _ _ R _ _ + (adic_valued.has_uniform_continuous_const_smul' R K v) + +@[simp] lemma coe_smul_adic_completion (r : R) (x : K) : + (↑(r • x) : v.adic_completion K) = r • (↑x : v.adic_completion K) := +@uniform_space.completion.coe_smul R K v.adic_valued.to_uniform_space _ _ r x + +instance : algebra K (v.adic_completion K) := +@uniform_space.completion.algebra' K _ v.adic_valued.to_uniform_space _ _ + +lemma algebra_map_adic_completion' : + ⇑(algebra_map R $ v.adic_completion K) = coe ∘ algebra_map R K := +rfl + +lemma algebra_map_adic_completion : + ⇑(algebra_map K $ v.adic_completion K) = coe := +rfl + +instance : is_scalar_tower R K (v.adic_completion K) := +@uniform_space.completion.is_scalar_tower R K K v.adic_valued.to_uniform_space _ _ _ + (adic_valued.has_uniform_continuous_const_smul' R K v) _ _ + +instance : algebra R (v.adic_completion_integers K) := +{ smul := λ r x, ⟨r • (x : v.adic_completion K), begin + have h : ((algebra_map R (adic_completion K v)) r) = (coe $ algebra_map R K r) := rfl, + rw algebra.smul_def, + refine valuation_subring.mul_mem _ _ _ _ x.2, + rw [mem_adic_completion_integers, h, valued.valued_completion_apply], + exact v.valuation_le_one _, + end⟩, + to_fun := λ r, ⟨coe $ algebra_map R K r, by simpa only [mem_adic_completion_integers, + valued.valued_completion_apply] using v.valuation_le_one _⟩, + map_one' := by simp only [map_one]; refl, + map_mul' := λ x y, + begin + ext, + simp_rw [ring_hom.map_mul, subring.coe_mul, subtype.coe_mk, uniform_space.completion.coe_mul], + end, + map_zero' := by simp only [map_zero]; refl, + map_add' := λ x y, + begin + ext, + simp_rw [ring_hom.map_add, subring.coe_add, subtype.coe_mk, uniform_space.completion.coe_add], + end, + commutes' := λ r x, by rw mul_comm, + smul_def' := λ r x, begin + ext, + simp only [subring.coe_mul, set_like.coe_mk, algebra.smul_def], + refl, + end } + +@[simp] lemma coe_smul_adic_completion_integers (r : R) (x : v.adic_completion_integers K) : + (↑(r • x) : v.adic_completion K) = r • (x : v.adic_completion K) := +rfl + +instance : no_zero_smul_divisors R (v.adic_completion_integers K) := +{ eq_zero_or_eq_zero_of_smul_eq_zero := λ c x hcx, + begin + rw [algebra.smul_def, mul_eq_zero] at hcx, + refine hcx.imp_left (λ hc, _), + letI : uniform_space K := v.adic_valued.to_uniform_space, + rw ← map_zero (algebra_map R (v.adic_completion_integers K)) at hc, + exact (is_fraction_ring.injective R K + (uniform_space.completion.coe_injective K (subtype.ext_iff.mp hc))) + end } + +instance adic_completion.is_scalar_tower' : + is_scalar_tower R (v.adic_completion_integers K) (v.adic_completion K) := +{ smul_assoc := λ x y z, by {simp only [algebra.smul_def], apply mul_assoc, }} + +end algebra_instances + end is_dedekind_domain.height_one_spectrum diff --git a/src/ring_theory/dedekind_domain/finite_adele_ring.lean b/src/ring_theory/dedekind_domain/finite_adele_ring.lean new file mode 100644 index 0000000000000..48c95752f18df --- /dev/null +++ b/src/ring_theory/dedekind_domain/finite_adele_ring.lean @@ -0,0 +1,228 @@ +/- +Copyright (c) 2023 María Inés de Frutos-Fernández. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: María Inés de Frutos-Fernández +-/ +import ring_theory.dedekind_domain.adic_valuation +import topology.algebra.uniform_ring + + +/-! +# The finite adèle ring of a Dedekind domain +We define the ring of finite adèles of a Dedekind domain `R`. + +## Main definitions +- `dedekind_domain.finite_integral_adeles` : product of `adic_completion_integers`, where `v` + runs over all maximal ideals of `R`. +- `dedekind_domain.prod_adic_completions` : the product of `adic_completion`, where `v` runs over + all maximal ideals of `R`. +- `dedekind_domain.finite_adele_ring` : The finite adèle ring of `R`, defined as the + restricted product `Π'_v K_v`. + +## Implementation notes +We are only interested on Dedekind domains of Krull dimension 1 (i.e., not fields). If `R` is a +field, its finite adèle ring is just defined to be the trivial ring. + +## References +* [J.W.S. Cassels, A. Frölich, *Algebraic Number Theory*][cassels1967algebraic] + +## Tags +finite adèle ring, dedekind domain +-/ + +noncomputable theory +open function set is_dedekind_domain is_dedekind_domain.height_one_spectrum + +namespace dedekind_domain + +variables (R K : Type*) [comm_ring R] [is_domain R] [is_dedekind_domain R] [field K] + [algebra R K] [is_fraction_ring R K] (v : height_one_spectrum R) + +/-- The product of all `adic_completion_integers`, where `v` runs over the maximal ideals of `R`. -/ +@[derive [comm_ring, topological_space, inhabited]] +def finite_integral_adeles : Type* := Π (v : height_one_spectrum R), v.adic_completion_integers K + +local notation `R_hat` := finite_integral_adeles + +/-- The product of all `adic_completion`, where `v` runs over the maximal ideals of `R`. -/ +@[derive [non_unital_non_assoc_ring, topological_space, topological_ring, comm_ring, inhabited]] +def prod_adic_completions := Π (v : height_one_spectrum R), v.adic_completion K + +local notation `K_hat` := prod_adic_completions + +namespace finite_integral_adeles + +noncomputable! instance : has_coe (R_hat R K) (K_hat R K) := { coe := λ x v, x v } + +lemma coe_apply (x : R_hat R K) (v : height_one_spectrum R) : (x : K_hat R K) v = ↑(x v) := rfl + +/-- The inclusion of `R_hat` in `K_hat` as a homomorphism of additive monoids. -/ +@[simps] def coe.add_monoid_hom : add_monoid_hom (R_hat R K) (K_hat R K) := +{ to_fun := coe, + map_zero' := rfl, + map_add' := λ x y, by { ext v, simp only [coe_apply, pi.add_apply, subring.coe_add] }} + +/-- The inclusion of `R_hat` in `K_hat` as a ring homomorphism. -/ +@[simps] def coe.ring_hom : ring_hom (R_hat R K) (K_hat R K) := +{ to_fun := coe, + map_one' := rfl, + map_mul' := λ x y, by {ext p, simp only [pi.mul_apply, subring.coe_mul], refl }, + ..coe.add_monoid_hom R K } + +end finite_integral_adeles + +section algebra_instances + +instance : algebra K (K_hat R K) := +(by apply_instance : algebra K $ Π v : height_one_spectrum R, v.adic_completion K) + +instance prod_adic_completions.algebra' : algebra R (K_hat R K) := +(by apply_instance : algebra R $ Π v : height_one_spectrum R, v.adic_completion K) + +instance : is_scalar_tower R K (K_hat R K) := +(by apply_instance : is_scalar_tower R K $ Π v : height_one_spectrum R, v.adic_completion K) + +instance : algebra R (R_hat R K) := +(by apply_instance : algebra R $ Π v : height_one_spectrum R, v.adic_completion_integers K) + +instance prod_adic_completions.algebra_completions : algebra (R_hat R K) (K_hat R K) := +(finite_integral_adeles.coe.ring_hom R K).to_algebra + +instance prod_adic_completions.is_scalar_tower_completions : + is_scalar_tower R (R_hat R K) (K_hat R K) := +(by apply_instance : is_scalar_tower R (Π v : height_one_spectrum R, v.adic_completion_integers K) $ + Π v : height_one_spectrum R, v.adic_completion K) + +end algebra_instances + +namespace finite_integral_adeles + +/-- The inclusion of `R_hat` in `K_hat` as an algebra homomorphism. -/ +def coe.alg_hom : alg_hom R (R_hat R K) (K_hat R K) := +{ to_fun := coe, + commutes' := λ r, rfl, + ..coe.ring_hom R K } + +lemma coe.alg_hom_apply (x : R_hat R K) (v : height_one_spectrum R) : + (coe.alg_hom R K) x v = x v := rfl + +end finite_integral_adeles + +/-! ### The finite adèle ring of a Dedekind domain +We define the finite adèle ring of `R` as the restricted product over all maximal ideals `v` of `R` +of `adic_completion` with respect to `adic_completion_integers`. We prove that it is a commutative +ring. TODO: show that it is a topological ring with the restricted product topology. -/ + +namespace prod_adic_completions + +variables {R K} + +/-- An element `x : K_hat R K` is a finite adèle if for all but finitely many height one ideals + `v`, the component `x v` is a `v`-adic integer. -/ +def is_finite_adele (x : K_hat R K) := +∀ᶠ v : height_one_spectrum R in filter.cofinite, x v ∈ v.adic_completion_integers K + +namespace is_finite_adele + +/-- The sum of two finite adèles is a finite adèle. -/ +lemma add {x y : K_hat R K} (hx : x.is_finite_adele) (hy : y.is_finite_adele) : + (x + y).is_finite_adele := +begin + rw [is_finite_adele, filter.eventually_cofinite] at hx hy ⊢, + have h_subset : {v : height_one_spectrum R | ¬ (x + y) v ∈ (v.adic_completion_integers K)} ⊆ + {v : height_one_spectrum R | ¬ x v ∈ (v.adic_completion_integers K)} ∪ + {v : height_one_spectrum R | ¬ y v ∈ (v.adic_completion_integers K)}, + { intros v hv, + rw [mem_union, mem_set_of_eq, mem_set_of_eq], + rw mem_set_of_eq at hv, + contrapose! hv, + rw [mem_adic_completion_integers, mem_adic_completion_integers, ← max_le_iff] at hv, + rw [mem_adic_completion_integers, pi.add_apply], + exact le_trans (valued.v.map_add_le_max' (x v) (y v)) hv }, + exact (hx.union hy).subset h_subset, +end + +/-- The tuple `(0)_v` is a finite adèle. -/ +lemma zero : (0 : K_hat R K).is_finite_adele := +begin + rw [is_finite_adele, filter.eventually_cofinite], + have h_empty : {v : height_one_spectrum R | + ¬ ((0 : v.adic_completion K) ∈ v.adic_completion_integers K)} = ∅, + { ext v, rw [mem_empty_iff_false, iff_false], intro hv, + rw mem_set_of_eq at hv, apply hv, rw mem_adic_completion_integers, + have h_zero : (valued.v (0 : v.adic_completion K) : (with_zero(multiplicative ℤ))) = 0 := + valued.v.map_zero', + rw h_zero, exact zero_le_one' _ }, + simp_rw [pi.zero_apply, h_empty], + exact finite_empty, +end + +/-- The negative of a finite adèle is a finite adèle. -/ +lemma neg {x : K_hat R K} (hx : x.is_finite_adele) : (-x).is_finite_adele := +begin + rw is_finite_adele at hx ⊢, + have h : ∀ (v : height_one_spectrum R), (-x v ∈ v.adic_completion_integers K) ↔ + (x v ∈ v.adic_completion_integers K), + { intro v, + rw [mem_adic_completion_integers, mem_adic_completion_integers, valuation.map_neg], }, + simpa only [pi.neg_apply, h] using hx, +end + +/-- The product of two finite adèles is a finite adèle. -/ +lemma mul {x y : K_hat R K} (hx : x.is_finite_adele) (hy : y.is_finite_adele) : + (x * y).is_finite_adele := +begin + rw [is_finite_adele, filter.eventually_cofinite] at hx hy ⊢, + have h_subset : {v : height_one_spectrum R | ¬ (x * y) v ∈ (v.adic_completion_integers K)} ⊆ + {v : height_one_spectrum R | ¬ x v ∈ (v.adic_completion_integers K)} ∪ + {v : height_one_spectrum R | ¬ y v ∈ (v.adic_completion_integers K)}, + { intros v hv, + rw [mem_union, mem_set_of_eq, mem_set_of_eq], + rw mem_set_of_eq at hv, + contrapose! hv, + rw [mem_adic_completion_integers, mem_adic_completion_integers] at hv, + have h_mul : valued.v (x v * y v) = (valued.v (x v)) * (valued.v (y v)) := + (valued.v).map_mul' (x v) (y v), + rw [mem_adic_completion_integers, pi.mul_apply, h_mul], + exact @mul_le_one' (with_zero (multiplicative ℤ)) _ _ + (ordered_comm_monoid.to_covariant_class_left _) _ _ hv.left hv.right }, + exact (hx.union hy).subset h_subset, +end + +/-- The tuple `(1)_v` is a finite adèle. -/ +lemma one : (1 : K_hat R K).is_finite_adele := +begin + rw [is_finite_adele, filter.eventually_cofinite], + have h_empty : {v : height_one_spectrum R | + ¬ ((1 : v.adic_completion K) ∈ v.adic_completion_integers K)} = ∅, + { ext v, rw [mem_empty_iff_false, iff_false], intro hv, + rw mem_set_of_eq at hv, apply hv, rw mem_adic_completion_integers, + exact le_of_eq valued.v.map_one' }, + simp_rw [pi.one_apply, h_empty], + exact finite_empty, +end + +end is_finite_adele + +end prod_adic_completions + +open prod_adic_completions.is_finite_adele + +variables (R K) +/-- The finite adèle ring of `R` is the restricted product over all maximal ideals `v` of `R` +of `adic_completion` with respect to `adic_completion_integers`. -/ +noncomputable! def finite_adele_ring : subring (K_hat R K) := +{ carrier := { x : K_hat R K | x.is_finite_adele }, + mul_mem' := λ _ _ hx hy, mul hx hy, + one_mem' := one, + add_mem' := λ _ _ hx hy, add hx hy, + zero_mem' := zero, + neg_mem' := λ _ hx, neg hx, } + +variables {R K} + +@[simp] lemma mem_finite_adele_ring_iff (x : K_hat R K) : + x ∈ finite_adele_ring R K ↔ x.is_finite_adele := +iff.rfl + +end dedekind_domain diff --git a/src/topology/algebra/valued_field.lean b/src/topology/algebra/valued_field.lean index 849a9b8268ac1..00eeec0472942 100644 --- a/src/topology/algebra/valued_field.lean +++ b/src/topology/algebra/valued_field.lean @@ -349,4 +349,7 @@ noncomputable instance valued_completion : valued (hat K) Γ₀ := exact (has_basis_nhds_zero K Γ₀).has_basis_of_dense_inducing completion.dense_inducing_coe, end } +@[simp, norm_cast] lemma valued_completion_apply (x : K) : valued.v (x : hat K) = v x := +extension_extends x + end valued From 5be98b9519379fd4434bb72c5e4120b870b8ff94 Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Fri, 24 Feb 2023 15:00:42 +0000 Subject: [PATCH 0523/1291] feat(geometry/manifold/vector_bundle): `smooth_fiberwise_linear` groupoid (#17302) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Let `B` be a smooth manifold and `F` a normed space. We build the groupoid of local homeomorphisms of `B × F` which are "smooth and fibrewise linear": that is, they are of the form `λ x, (x.1, φ x.1 x.2)` for some function `φ` from `B` to the automorphisms of `F`, which is smooth and has smooth inverse on the proposed domain. Membership in this groupoid is the natural property characterizing the transition functions of a smooth vector bundle over `B` modelled on `F`. Co-authored-by: Floris van Doorn Co-authored-by: Floris van Doorn --- src/geometry/manifold/cont_mdiff.lean | 20 +- src/geometry/manifold/vector_bundle.lean | 277 +++++++++++++++++++++++ 2 files changed, 292 insertions(+), 5 deletions(-) create mode 100644 src/geometry/manifold/vector_bundle.lean diff --git a/src/geometry/manifold/cont_mdiff.lean b/src/geometry/manifold/cont_mdiff.lean index 7538d4211da58..384933add370c 100644 --- a/src/geometry/manifold/cont_mdiff.lean +++ b/src/geometry/manifold/cont_mdiff.lean @@ -1624,14 +1624,24 @@ lemma continuous_linear_map.cont_mdiff (L : E →L[𝕜] F) : L.cont_diff.cont_mdiff -- the following proof takes very long to elaborate in pure term mode +lemma cont_mdiff_within_at.clm_comp {g : M → F →L[𝕜] F''} {f : M → F' →L[𝕜] F} {s : set M} {x : M} + (hg : cont_mdiff_within_at I 𝓘(𝕜, F →L[𝕜] F'') n g s x) + (hf : cont_mdiff_within_at I 𝓘(𝕜, F' →L[𝕜] F) n f s x) : + cont_mdiff_within_at I 𝓘(𝕜, F' →L[𝕜] F'') n (λ x, (g x).comp (f x)) s x := +@cont_diff_within_at.comp_cont_mdiff_within_at _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ + (λ x : (F →L[𝕜] F'') × (F' →L[𝕜] F), x.1.comp x.2) (λ x, (g x, f x)) s _ x + (by { apply cont_diff.cont_diff_at, exact cont_diff_fst.clm_comp cont_diff_snd }) + (hg.prod_mk_space hf) (by simp_rw [preimage_univ, subset_univ]) + lemma cont_mdiff_at.clm_comp {g : M → F →L[𝕜] F''} {f : M → F' →L[𝕜] F} {x : M} (hg : cont_mdiff_at I 𝓘(𝕜, F →L[𝕜] F'') n g x) (hf : cont_mdiff_at I 𝓘(𝕜, F' →L[𝕜] F) n f x) : cont_mdiff_at I 𝓘(𝕜, F' →L[𝕜] F'') n (λ x, (g x).comp (f x)) x := -@cont_diff_at.comp_cont_mdiff_at _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ - (λ x : (F →L[𝕜] F'') × (F' →L[𝕜] F), x.1.comp x.2) (λ x, (g x, f x)) x - (by { apply cont_diff.cont_diff_at, apply is_bounded_bilinear_map.cont_diff, - exact is_bounded_bilinear_map_comp }) -- todo: simplify after #16946 - (hg.prod_mk_space hf) +(hg.cont_mdiff_within_at.clm_comp hf.cont_mdiff_within_at).cont_mdiff_at univ_mem + +lemma cont_mdiff_on.clm_comp {g : M → F →L[𝕜] F''} {f : M → F' →L[𝕜] F} {s : set M} + (hg : cont_mdiff_on I 𝓘(𝕜, F →L[𝕜] F'') n g s) (hf : cont_mdiff_on I 𝓘(𝕜, F' →L[𝕜] F) n f s) : + cont_mdiff_on I 𝓘(𝕜, F' →L[𝕜] F'') n (λ x, (g x).comp (f x)) s := +λ x hx, (hg x hx).clm_comp (hf x hx) lemma cont_mdiff.clm_comp {g : M → F →L[𝕜] F''} {f : M → F' →L[𝕜] F} (hg : cont_mdiff I 𝓘(𝕜, F →L[𝕜] F'') n g) (hf : cont_mdiff I 𝓘(𝕜, F' →L[𝕜] F) n f) : diff --git a/src/geometry/manifold/vector_bundle.lean b/src/geometry/manifold/vector_bundle.lean new file mode 100644 index 0000000000000..6398718dc840f --- /dev/null +++ b/src/geometry/manifold/vector_bundle.lean @@ -0,0 +1,277 @@ +/- +Copyright (c) 2022 Floris van Doorn, Heather Macbeth. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Floris van Doorn, Heather Macbeth +-/ +import geometry.manifold.cont_mdiff + +/-! # Smooth vector bundles + +This file will eventually contain the definition of a smooth vector bundle. For now, it contains +preliminaries regarding an associated `structure_groupoid`, the groupoid of +`smooth_fibrewise_linear` functions. When a (topological) vector bundle is smooth, then the +composition of charts associated to the vector bundle belong to this groupoid. +-/ + +noncomputable theory + +open set topological_space +open_locale manifold topology + +/-! ### The groupoid of smooth, fibrewise-linear maps -/ + +variables {𝕜 B F : Type*} [topological_space B] +variables [nontrivially_normed_field 𝕜] [normed_add_comm_group F] [normed_space 𝕜 F] + +namespace fiberwise_linear + +variables {φ φ' : B → F ≃L[𝕜] F} {U U' : set B} + +/-- For `B` a topological space and `F` a `𝕜`-normed space, a map from `U : set B` to `F ≃L[𝕜] F` +determines a local homeomorphism from `B × F` to itself by its action fibrewise. -/ +def local_homeomorph (φ : B → F ≃L[𝕜] F) (hU : is_open U) + (hφ : continuous_on (λ x, φ x : B → F →L[𝕜] F) U) + (h2φ : continuous_on (λ x, (φ x).symm : B → F →L[𝕜] F) U) : + local_homeomorph (B × F) (B × F) := +{ to_fun := λ x, (x.1, φ x.1 x.2), + inv_fun := λ x, (x.1, (φ x.1).symm x.2), + source := U ×ˢ univ, + target := U ×ˢ univ, + map_source' := λ x hx, mk_mem_prod hx.1 (mem_univ _), + map_target' := λ x hx, mk_mem_prod hx.1 (mem_univ _), + left_inv' := λ x _, prod.ext rfl (continuous_linear_equiv.symm_apply_apply _ _), + right_inv' := λ x _, prod.ext rfl (continuous_linear_equiv.apply_symm_apply _ _), + open_source := hU.prod is_open_univ, + open_target := hU.prod is_open_univ, + continuous_to_fun := begin + have : continuous_on (λ p : B × F, ((φ p.1 : F →L[𝕜] F), p.2)) (U ×ˢ univ), + { exact hφ.prod_map continuous_on_id }, + exact continuous_on_fst.prod (is_bounded_bilinear_map_apply.continuous.comp_continuous_on this), + end, + continuous_inv_fun := begin + have : continuous_on (λ p : B × F, (((φ p.1).symm : F →L[𝕜] F), p.2)) (U ×ˢ univ), + { exact h2φ.prod_map continuous_on_id }, + exact continuous_on_fst.prod (is_bounded_bilinear_map_apply.continuous.comp_continuous_on this), + end, } + +/-- Compute the composition of two local homeomorphisms induced by fiberwise linear +equivalences. -/ +lemma trans_local_homeomorph_apply + (hU : is_open U) + (hφ : continuous_on (λ x, φ x : B → F →L[𝕜] F) U) + (h2φ : continuous_on (λ x, (φ x).symm : B → F →L[𝕜] F) U) + (hU' : is_open U') + (hφ' : continuous_on (λ x, φ' x : B → F →L[𝕜] F) U') + (h2φ' : continuous_on (λ x, (φ' x).symm : B → F →L[𝕜] F) U') + (b : B) (v : F) : + (fiberwise_linear.local_homeomorph φ hU hφ h2φ ≫ₕ + fiberwise_linear.local_homeomorph φ' hU' hφ' h2φ') ⟨b, v⟩ = ⟨b, φ' b (φ b v)⟩ := +rfl + +/-- Compute the source of the composition of two local homeomorphisms induced by fiberwise linear +equivalences. -/ +lemma source_trans_local_homeomorph + (hU : is_open U) + (hφ : continuous_on (λ x, φ x : B → F →L[𝕜] F) U) + (h2φ : continuous_on (λ x, (φ x).symm : B → F →L[𝕜] F) U) + (hU' : is_open U') + (hφ' : continuous_on (λ x, φ' x : B → F →L[𝕜] F) U') + (h2φ' : continuous_on (λ x, (φ' x).symm : B → F →L[𝕜] F) U') : + (fiberwise_linear.local_homeomorph φ hU hφ h2φ ≫ₕ + fiberwise_linear.local_homeomorph φ' hU' hφ' h2φ').source = (U ∩ U') ×ˢ univ := +by { dsimp [fiberwise_linear.local_homeomorph], mfld_set_tac } + +/-- Compute the target of the composition of two local homeomorphisms induced by fiberwise linear +equivalences. -/ +lemma target_trans_local_homeomorph + (hU : is_open U) + (hφ : continuous_on (λ x, φ x : B → F →L[𝕜] F) U) + (h2φ : continuous_on (λ x, (φ x).symm : B → F →L[𝕜] F) U) + (hU' : is_open U') + (hφ' : continuous_on (λ x, φ' x : B → F →L[𝕜] F) U') + (h2φ' : continuous_on (λ x, (φ' x).symm : B → F →L[𝕜] F) U') : + (fiberwise_linear.local_homeomorph φ hU hφ h2φ ≫ₕ + fiberwise_linear.local_homeomorph φ' hU' hφ' h2φ').target = (U ∩ U') ×ˢ univ := +by { dsimp [fiberwise_linear.local_homeomorph], mfld_set_tac } + +end fiberwise_linear + +variables {EB : Type*} [normed_add_comm_group EB] [normed_space 𝕜 EB] + {HB : Type*} [topological_space HB] [charted_space HB B] {IB : model_with_corners 𝕜 EB HB} + +/-- Let `e` be a local homeomorphism of `B × F`. Suppose that at every point `p` in the source of +`e`, there is some neighbourhood `s` of `p` on which `e` is equal to a bi-smooth fibrewise linear +local homeomorphism. + +Then the source of `e` is of the form `U ×ˢ univ`, for some set `U` in `B`, and, at any point `x` in +`U`, admits a neighbourhood `u` of `x` such that `e` is equal on `u ×ˢ univ` to some bi-smooth +fibrewise linear local homeomorphism. -/ +lemma smooth_fibrewise_linear.locality_aux₁ (e : local_homeomorph (B × F) (B × F)) + (h : ∀ p ∈ e.source, ∃ s : set (B × F), is_open s ∧ p ∈ s ∧ + ∃ (φ : B → (F ≃L[𝕜] F)) (u : set B) (hu : is_open u) + (hφ : smooth_on IB 𝓘(𝕜, F →L[𝕜] F) (λ x, (φ x : F →L[𝕜] F)) u) + (h2φ : smooth_on IB 𝓘(𝕜, F →L[𝕜] F) (λ x, ((φ x).symm : F →L[𝕜] F)) u), + (e.restr s).eq_on_source + (fiberwise_linear.local_homeomorph φ hu hφ.continuous_on h2φ.continuous_on)) : + ∃ (U : set B) (hU : e.source = U ×ˢ univ), + ∀ x ∈ U, ∃ (φ : B → (F ≃L[𝕜] F)) (u : set B) (hu : is_open u) (huU : u ⊆ U) (hux : x ∈ u) + (hφ : smooth_on IB 𝓘(𝕜, F →L[𝕜] F) (λ x, (φ x : F →L[𝕜] F)) u) + (h2φ : smooth_on IB 𝓘(𝕜, F →L[𝕜] F) (λ x, ((φ x).symm : F →L[𝕜] F)) u), + (e.restr (u ×ˢ univ)).eq_on_source + (fiberwise_linear.local_homeomorph φ hu hφ.continuous_on h2φ.continuous_on) := +begin + rw set_coe.forall' at h, + choose s hs hsp φ u hu hφ h2φ heφ using h, + have hesu : ∀ p : e.source, e.source ∩ s p = u p ×ˢ univ, + { intros p, + rw ← e.restr_source' (s _) (hs _), + exact (heφ p).1 }, + have hu' : ∀ p : e.source, (p : B × F).fst ∈ u p, + { intros p, + have : (p : B × F) ∈ e.source ∩ s p := ⟨p.prop, hsp p⟩, + simpa only [hesu, mem_prod, mem_univ, and_true] using this }, + have heu : ∀ p : e.source, ∀ q : B × F, q.fst ∈ u p → q ∈ e.source, + { intros p q hq, + have : q ∈ u p ×ˢ (univ : set F) := ⟨hq, trivial⟩, + rw ← hesu p at this, + exact this.1 }, + have he : e.source = (prod.fst '' e.source) ×ˢ (univ : set F), + { apply has_subset.subset.antisymm, + { intros p hp, + exact ⟨⟨p, hp, rfl⟩, trivial⟩ }, + { rintros ⟨x, v⟩ ⟨⟨p, hp, rfl : p.fst = x⟩, -⟩, + exact heu ⟨p, hp⟩ (p.fst, v) (hu' ⟨p, hp⟩) } }, + refine ⟨prod.fst '' e.source, he, _⟩, + rintros x ⟨p, hp, rfl⟩, + refine ⟨φ ⟨p, hp⟩, u ⟨p, hp⟩, hu ⟨p, hp⟩, _, hu' _, hφ ⟨p, hp⟩, h2φ ⟨p, hp⟩, _⟩, + { intros y hy, refine ⟨(y, 0), heu ⟨p, hp⟩ ⟨_, _⟩ hy, rfl⟩ }, + { rw [← hesu, e.restr_source_inter], exact heφ ⟨p, hp⟩ }, +end + +/-- Let `e` be a local homeomorphism of `B × F` whose source is `U ×ˢ univ`, for some set `U` in +`B`, and which, at any point `x` in `U`, admits a neighbourhood `u` of `x` such that `e` is equal on +`u ×ˢ univ` to some bi-smooth fibrewise linear local homeomorphism. Then `e` itself is equal to +some bi-smooth fibrewise linear local homeomorphism. + +This is the key mathematical point of the `locality` condition in the construction of the +`structure_groupoid` of bi-smooth fibrewise linear local homeomorphisms. The proof is by gluing +together the various bi-smooth fibrewise linear local homeomorphism which exist locally. + +The `U` in the conclusion is the same `U` as in the hypothesis. We state it like this, because this +is exactly what we need for `smooth_fiberwise_linear`. -/ +lemma smooth_fibrewise_linear.locality_aux₂ (e : local_homeomorph (B × F) (B × F)) + (U : set B) (hU : e.source = U ×ˢ univ) + (h : ∀ x ∈ U, ∃ (φ : B → (F ≃L[𝕜] F)) (u : set B) (hu : is_open u) (hUu : u ⊆ U) (hux : x ∈ u) + (hφ : smooth_on IB 𝓘(𝕜, F →L[𝕜] F) (λ x, (φ x : F →L[𝕜] F)) u) + (h2φ : smooth_on IB 𝓘(𝕜, F →L[𝕜] F) (λ x, ((φ x).symm : F →L[𝕜] F)) u), + (e.restr (u ×ˢ univ)).eq_on_source + (fiberwise_linear.local_homeomorph φ hu hφ.continuous_on h2φ.continuous_on)) : + ∃ (Φ : B → (F ≃L[𝕜] F)) (U : set B) (hU₀ : is_open U) + (hΦ : smooth_on IB 𝓘(𝕜, F →L[𝕜] F) (λ x, (Φ x : F →L[𝕜] F)) U) + (h2Φ : smooth_on IB 𝓘(𝕜, F →L[𝕜] F) (λ x, ((Φ x).symm : F →L[𝕜] F)) U), + e.eq_on_source (fiberwise_linear.local_homeomorph Φ hU₀ hΦ.continuous_on h2Φ.continuous_on) := +begin + classical, + rw set_coe.forall' at h, + choose! φ u hu hUu hux hφ h2φ heφ using h, + have heuφ : ∀ x : U, eq_on e (λ q, (q.1, φ x q.1 q.2)) (u x ×ˢ univ), + { intros x p hp, + refine (heφ x).2 _, + rw (heφ x).1, + exact hp }, + have huφ : ∀ (x x' : U) (y : B) (hyx : y ∈ u x) (hyx' : y ∈ u x'), φ x y = φ x' y, + { intros p p' y hyp hyp', + ext v, + have h1 : e (y, v) = (y, φ p y v) := heuφ _ ⟨(id hyp : (y, v).fst ∈ u p), trivial⟩, + have h2 : e (y, v) = (y, φ p' y v) := heuφ _ ⟨(id hyp' : (y, v).fst ∈ u p'), trivial⟩, + exact congr_arg prod.snd (h1.symm.trans h2) }, + have hUu' : U = ⋃ i, u i, + { ext x, + rw mem_Union, + refine ⟨λ h, ⟨⟨x, h⟩, hux _⟩, _⟩, + rintros ⟨x, hx⟩, + exact hUu x hx }, + have hU' : is_open U, + { rw hUu', + apply is_open_Union hu }, + let Φ₀ : U → F ≃L[𝕜] F := Union_lift u (λ x, (φ x) ∘ coe) huφ U hUu'.le, + let Φ : B → F ≃L[𝕜] F := λ y, if hy : y ∈ U then Φ₀ ⟨y, hy⟩ else continuous_linear_equiv.refl 𝕜 F, + have hΦ : ∀ (y) (hy : y ∈ U), Φ y = Φ₀ ⟨y, hy⟩ := λ y hy, dif_pos hy, + have hΦφ : ∀ x : U, ∀ y ∈ u x, Φ y = φ x y, + { intros x y hyu, + refine (hΦ y (hUu x hyu)).trans _, + exact Union_lift_mk ⟨y, hyu⟩ _ }, + have hΦ : smooth_on IB 𝓘(𝕜, F →L[𝕜] F) (λ y, (Φ y : F →L[𝕜] F)) U, + { apply cont_mdiff_on_of_locally_cont_mdiff_on, + intros x hx, + refine ⟨u ⟨x, hx⟩, hu ⟨x, hx⟩, hux _, _⟩, + refine (cont_mdiff_on.congr (hφ ⟨x, hx⟩) _).mono (inter_subset_right _ _), + intros y hy, + rw hΦφ ⟨x, hx⟩ y hy }, + have h2Φ : smooth_on IB 𝓘(𝕜, F →L[𝕜] F) (λ y, ((Φ y).symm : F →L[𝕜] F)) U, + { apply cont_mdiff_on_of_locally_cont_mdiff_on, + intros x hx, + refine ⟨u ⟨x, hx⟩, hu ⟨x, hx⟩, hux _, _⟩, + refine (cont_mdiff_on.congr (h2φ ⟨x, hx⟩) _).mono (inter_subset_right _ _), + intros y hy, + rw hΦφ ⟨x, hx⟩ y hy }, + refine ⟨Φ, U, hU', hΦ, h2Φ, hU, λ p hp, _⟩, + rw [hU] at hp, + -- using rw on the next line seems to cause a timeout in kernel type-checking + refine (heuφ ⟨p.fst, hp.1⟩ ⟨hux _, hp.2⟩).trans _, + congrm (_, _), + rw hΦφ, + apply hux +end + +variables (F B IB) +/-- For `B` a manifold and `F` a normed space, the groupoid on `B × F` consisting of local +homeomorphisms which are bi-smooth and fibrewise linear, and induce the identity on `B`. +When a (topological) vector bundle is smooth, then the composition of charts associated +to the vector bundle belong to this groupoid. -/ +def smooth_fiberwise_linear : structure_groupoid (B × F) := +{ members := ⋃ (φ : B → F ≃L[𝕜] F) (U : set B) (hU : is_open U) + (hφ : smooth_on IB 𝓘(𝕜, F →L[𝕜] F) (λ x, φ x : B → F →L[𝕜] F) U) + (h2φ : smooth_on IB 𝓘(𝕜, F →L[𝕜] F) (λ x, (φ x).symm : B → F →L[𝕜] F) U), + {e | e.eq_on_source (fiberwise_linear.local_homeomorph φ hU hφ.continuous_on h2φ.continuous_on)}, + trans' := begin + simp_rw [mem_Union], + rintros e e' ⟨φ, U, hU, hφ, h2φ, heφ⟩ ⟨φ', U', hU', hφ', h2φ', heφ'⟩, + refine ⟨λ b, (φ b).trans (φ' b), _, hU.inter hU', _, _, setoid.trans (heφ.trans' heφ') ⟨_, _⟩⟩, + { show smooth_on IB 𝓘(𝕜, F →L[𝕜] F) + (λ (x : B), (φ' x).to_continuous_linear_map ∘L (φ x).to_continuous_linear_map) (U ∩ U'), + exact (hφ'.mono $ inter_subset_right _ _).clm_comp (hφ.mono $ inter_subset_left _ _) }, + { show smooth_on IB 𝓘(𝕜, F →L[𝕜] F) (λ (x : B), + (φ x).symm.to_continuous_linear_map ∘L (φ' x).symm.to_continuous_linear_map) (U ∩ U'), + exact (h2φ.mono $ inter_subset_left _ _).clm_comp (h2φ'.mono $ inter_subset_right _ _) }, + { apply fiberwise_linear.source_trans_local_homeomorph }, + { rintros ⟨b, v⟩ hb, apply fiberwise_linear.trans_local_homeomorph_apply } + end, + symm' := begin + simp_rw [mem_Union], + rintros e ⟨φ, U, hU, hφ, h2φ, heφ⟩, + refine ⟨λ b, (φ b).symm, U, hU, h2φ, _, heφ.symm'⟩, + simp_rw continuous_linear_equiv.symm_symm, + exact hφ + end, + id_mem' := begin + simp_rw [mem_Union], + refine ⟨λ b, continuous_linear_equiv.refl 𝕜 F, univ, is_open_univ, _, _, ⟨_, λ b hb, _⟩⟩, + { apply cont_mdiff_on_const }, + { apply cont_mdiff_on_const }, + { simp only [fiberwise_linear.local_homeomorph, local_homeomorph.refl_local_equiv, + local_equiv.refl_source, univ_prod_univ] }, + { simp only [fiberwise_linear.local_homeomorph, local_homeomorph.refl_apply, prod.mk.eta, + id.def, continuous_linear_equiv.coe_refl', local_homeomorph.mk_coe, local_equiv.coe_mk] }, + end, + locality' := begin -- the hard work has been extracted to `locality_aux₁` and `locality_aux₂` + simp_rw [mem_Union], + intros e he, + obtain ⟨U, hU, h⟩ := smooth_fibrewise_linear.locality_aux₁ e he, + exact smooth_fibrewise_linear.locality_aux₂ e U hU h, + end, + eq_on_source' := begin + simp_rw [mem_Union], + rintros e e' ⟨φ, U, hU, hφ, h2φ, heφ⟩ hee', + exact ⟨φ, U, hU, hφ, h2φ, setoid.trans hee' heφ⟩, + end } From afdb4fa3b32d41106a4a09b371ce549ad7958abd Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Fri, 24 Feb 2023 15:00:44 +0000 Subject: [PATCH 0524/1291] feat(*/with_top): add lemmas about `with_top`/`with_bot` (#18487) Backport leanprover-community/mathlib4#2406 --- src/algebra/big_operators/order.lean | 42 ++++------- src/algebra/order/monoid/with_top.lean | 6 +- src/algebra/order/ring/with_top.lean | 70 ++++++++++--------- src/algebra/order/sub/with_top.lean | 5 ++ src/analysis/bounded_variation.lean | 7 +- src/analysis/mean_inequalities_pow.lean | 2 +- src/analysis/p_series.lean | 2 +- src/data/real/ennreal.lean | 27 +++---- src/data/real/ereal.lean | 3 +- src/measure_theory/function/l1_space.lean | 3 +- src/measure_theory/function/lp_space.lean | 19 ++--- .../integral/mean_inequalities.lean | 3 +- src/order/with_bot.lean | 32 +++++++++ src/topology/instances/ennreal.lean | 4 +- .../metric_space/hausdorff_distance.lean | 5 +- 15 files changed, 126 insertions(+), 104 deletions(-) diff --git a/src/algebra/big_operators/order.lean b/src/algebra/big_operators/order.lean index ac17f0708b860..25dd5b41c238d 100644 --- a/src/algebra/big_operators/order.lean +++ b/src/algebra/big_operators/order.lean @@ -602,34 +602,26 @@ namespace with_top open finset /-- A product of finite numbers is still finite -/ -lemma prod_lt_top [canonically_ordered_comm_semiring R] [nontrivial R] [decidable_eq R] - {s : finset ι} {f : ι → with_top R} (h : ∀ i ∈ s, f i ≠ ⊤) : +lemma prod_lt_top [comm_monoid_with_zero R] [no_zero_divisors R] [nontrivial R] [decidable_eq R] + [has_lt R] {s : finset ι} {f : ι → with_top R} (h : ∀ i ∈ s, f i ≠ ⊤) : ∏ i in s, f i < ⊤ := -prod_induction f (λ a, a < ⊤) (λ a b h₁ h₂, mul_lt_top h₁.ne h₂.ne) (coe_lt_top 1) $ - λ a ha, lt_top_iff_ne_top.2 (h a ha) - -/-- A sum of finite numbers is still finite -/ -lemma sum_lt_top [ordered_add_comm_monoid M] {s : finset ι} {f : ι → with_top M} - (h : ∀ i ∈ s, f i ≠ ⊤) : (∑ i in s, f i) < ⊤ := -sum_induction f (λ a, a < ⊤) (λ a b h₁ h₂, add_lt_top.2 ⟨h₁, h₂⟩) zero_lt_top $ - λ i hi, lt_top_iff_ne_top.2 (h i hi) +prod_induction f (λ a, a < ⊤) (λ a b h₁ h₂, mul_lt_top' h₁ h₂) (coe_lt_top 1) $ + λ a ha, with_top.lt_top_iff_ne_top.2 (h a ha) /-- A sum of numbers is infinite iff one of them is infinite -/ -lemma sum_eq_top_iff [ordered_add_comm_monoid M] {s : finset ι} {f : ι → with_top M} : +lemma sum_eq_top_iff [add_comm_monoid M] {s : finset ι} {f : ι → with_top M} : ∑ i in s, f i = ⊤ ↔ ∃ i ∈ s, f i = ⊤ := -begin - classical, - split, - { contrapose!, - exact λ h, (sum_lt_top $ λ i hi, (h i hi)).ne }, - { rintro ⟨i, his, hi⟩, - rw [sum_eq_add_sum_diff_singleton his, hi, top_add] } -end +by induction s using finset.cons_induction; simp [*, or_and_distrib_right, exists_or_distrib] /-- A sum of finite numbers is still finite -/ -lemma sum_lt_top_iff [ordered_add_comm_monoid M] {s : finset ι} {f : ι → with_top M} : +lemma sum_lt_top_iff [add_comm_monoid M] [has_lt M] {s : finset ι} {f : ι → with_top M} : ∑ i in s, f i < ⊤ ↔ ∀ i ∈ s, f i < ⊤ := -by simp only [lt_top_iff_ne_top, ne.def, sum_eq_top_iff, not_exists] +by simp only [with_top.lt_top_iff_ne_top, ne.def, sum_eq_top_iff, not_exists] + +/-- A sum of finite numbers is still finite -/ +lemma sum_lt_top [add_comm_monoid M] [has_lt M] {s : finset ι} {f : ι → with_top M} + (h : ∀ i ∈ s, f i ≠ ⊤) : (∑ i in s, f i) < ⊤ := +sum_lt_top_iff.2 $ λ i hi, with_top.lt_top_iff_ne_top.2 (h i hi) end with_top @@ -640,13 +632,7 @@ variables {S : Type*} lemma absolute_value.sum_le [semiring R] [ordered_semiring S] (abv : absolute_value R S) (s : finset ι) (f : ι → R) : abv (∑ i in s, f i) ≤ ∑ i in s, abv (f i) := -begin - letI := classical.dec_eq ι, - refine finset.induction_on s _ (λ i s hi ih, _), - { simp }, - { simp only [finset.sum_insert hi], - exact (abv.add_le _ _).trans (add_le_add le_rfl ih) }, -end +finset.le_sum_of_subadditive abv (map_zero _) abv.add_le _ _ lemma is_absolute_value.abv_sum [semiring R] [ordered_semiring S] (abv : R → S) [is_absolute_value abv] (f : ι → R) (s : finset ι) : diff --git a/src/algebra/order/monoid/with_top.lean b/src/algebra/order/monoid/with_top.lean index b80bfa8232f70..3648e90254f39 100644 --- a/src/algebra/order/monoid/with_top.lean +++ b/src/algebra/order/monoid/with_top.lean @@ -74,8 +74,8 @@ by cases a; cases b; simp [none_eq_top, some_eq_coe, ←with_top.coe_add] lemma add_ne_top : a + b ≠ ⊤ ↔ a ≠ ⊤ ∧ b ≠ ⊤ := add_eq_top.not.trans not_or_distrib -lemma add_lt_top [partial_order α] {a b : with_top α} : a + b < ⊤ ↔ a < ⊤ ∧ b < ⊤ := -by simp_rw [lt_top_iff_ne_top, add_ne_top] +lemma add_lt_top [has_lt α] {a b : with_top α} : a + b < ⊤ ↔ a < ⊤ ∧ b < ⊤ := +by simp_rw [with_top.lt_top_iff_ne_top, add_ne_top] lemma add_eq_coe : ∀ {a b : with_top α} {c : α}, a + b = c ↔ ∃ (a' b' : α), ↑a' = a ∧ ↑b' = b ∧ a' + b' = c @@ -380,7 +380,7 @@ lemma coe_bit1 [has_one α] {a : α} : ((bit1 a : α) : with_bot α) = bit1 a := @[simp] lemma add_eq_bot : a + b = ⊥ ↔ a = ⊥ ∨ b = ⊥ := with_top.add_eq_top lemma add_ne_bot : a + b ≠ ⊥ ↔ a ≠ ⊥ ∧ b ≠ ⊥ := with_top.add_ne_top -lemma bot_lt_add [partial_order α] {a b : with_bot α} : ⊥ < a + b ↔ ⊥ < a ∧ ⊥ < b := +lemma bot_lt_add [has_lt α] {a b : with_bot α} : ⊥ < a + b ↔ ⊥ < a ∧ ⊥ < b := @with_top.add_lt_top αᵒᵈ _ _ _ _ lemma add_eq_coe : a + b = x ↔ ∃ (a' b' : α), ↑a' = a ∧ ↑b' = b ∧ a' + b' = x := with_top.add_eq_coe diff --git a/src/algebra/order/ring/with_top.lean b/src/algebra/order/ring/with_top.lean index 3cc33489d4e66..d39db797bd686 100644 --- a/src/algebra/order/ring/with_top.lean +++ b/src/algebra/order/ring/with_top.lean @@ -36,15 +36,36 @@ instance : mul_zero_class (with_top α) := lemma mul_def {a b : with_top α} : a * b = if a = 0 ∨ b = 0 then 0 else option.map₂ (*) a b := rfl -@[simp] lemma mul_top {a : with_top α} (h : a ≠ 0) : a * ⊤ = ⊤ := -by cases a; simp [mul_def, h]; refl +lemma mul_top' {a : with_top α} : a * ⊤ = if a = 0 then 0 else ⊤ := +by induction a using with_top.rec_top_coe; simp [mul_def]; refl -@[simp] lemma top_mul {a : with_top α} (h : a ≠ 0) : ⊤ * a = ⊤ := -by cases a; simp [mul_def, h]; refl +@[simp] lemma mul_top {a : with_top α} (h : a ≠ 0) : a * ⊤ = ⊤ := by rw [mul_top', if_neg h] + +lemma top_mul' {a : with_top α} : ⊤ * a = if a = 0 then 0 else ⊤ := +by induction a using with_top.rec_top_coe; simp [mul_def]; refl + +@[simp] lemma top_mul {a : with_top α} (h : a ≠ 0) : ⊤ * a = ⊤ := by rw [top_mul', if_neg h] @[simp] lemma top_mul_top : (⊤ * ⊤ : with_top α) = ⊤ := top_mul top_ne_zero +theorem mul_eq_top_iff {a b : with_top α} : a * b = ⊤ ↔ a ≠ 0 ∧ b = ⊤ ∨ a = ⊤ ∧ b ≠ 0 := +begin + rw [mul_def, ite_eq_iff, ← none_eq_top, option.map₂_eq_none_iff], + have ha : a = 0 → a ≠ none := λ h, h.symm ▸ zero_ne_top, + have hb : b = 0 → b ≠ none := λ h, h.symm ▸ zero_ne_top, + tauto +end + +theorem mul_lt_top' [has_lt α] {a b : with_top α} (ha : a < ⊤) (hb : b < ⊤) : a * b < ⊤ := +begin + rw [with_top.lt_top_iff_ne_top] at *, + simp only [ne.def, mul_eq_top_iff, *, and_false, false_and, false_or, not_false_iff] +end + +theorem mul_lt_top [has_lt α] {a b : with_top α} (ha : a ≠ ⊤) (hb : b ≠ ⊤) : a * b < ⊤ := + mul_lt_top' (with_top.lt_top_iff_ne_top.2 ha) (with_top.lt_top_iff_ne_top.2 hb) + instance [no_zero_divisors α] : no_zero_divisors (with_top α) := begin refine ⟨λ a b h₁, decidable.by_contradiction $ λ h₂, _⟩, @@ -59,7 +80,7 @@ section mul_zero_class variables [mul_zero_class α] -@[norm_cast] lemma coe_mul {a b : α} : (↑(a * b) : with_top α) = a * b := +@[simp, norm_cast] lemma coe_mul {a b : α} : (↑(a * b) : with_top α) = a * b := decidable.by_cases (assume : a = 0, by simp [this]) $ assume ha, decidable.by_cases (assume : b = 0, by simp [this]) $ assume hb, by { simp [*, mul_def] } @@ -69,22 +90,6 @@ lemma mul_coe {b : α} (hb : b ≠ 0) : ∀{a : with_top α}, a * b = a.bind (λ by simp [hb] | (some a) := show ↑a * ↑b = ↑(a * b), from coe_mul.symm -@[simp] lemma mul_eq_top_iff {a b : with_top α} : a * b = ⊤ ↔ (a ≠ 0 ∧ b = ⊤) ∨ (a = ⊤ ∧ b ≠ 0) := -begin - cases a; cases b; simp only [none_eq_top, some_eq_coe], - { simp [← coe_mul] }, - { by_cases hb : b = 0; simp [hb] }, - { by_cases ha : a = 0; simp [ha] }, - { simp [← coe_mul] } -end - -lemma mul_lt_top [preorder α] {a b : with_top α} (ha : a ≠ ⊤) (hb : b ≠ ⊤) : a * b < ⊤ := -begin - lift a to α using ha, - lift b to α using hb, - simp only [← coe_mul, coe_lt_top] -end - @[simp] lemma untop'_zero_mul (a b : with_top α) : (a * b).untop' 0 = a.untop' 0 * b.untop' 0 := begin by_cases ha : a = 0, { rw [ha, zero_mul, ← coe_zero, untop'_coe, zero_mul] }, @@ -127,7 +132,7 @@ instance [mul_zero_one_class α] [nontrivial α] : mul_zero_one_class (with_top induction y using with_top.rec_top_coe, { have : (f x : with_top S) ≠ 0, by simpa [hf.eq_iff' (map_zero f)] using hx, simp [hx, this] }, - simp [← coe_mul] + simp only [← coe_mul, map_coe, map_mul] end, .. f.to_zero_hom.with_top_map, .. f.to_monoid_hom.to_one_hom.with_top_map } @@ -163,7 +168,7 @@ begin induction c using with_top.rec_top_coe, { by_cases ha : a = 0; simp [ha] }, { by_cases hc : c = 0, { simp [hc] }, - simp [mul_coe hc], cases a; cases b, + simp only [mul_coe hc], cases a; cases b, repeat { refl <|> exact congr_arg some (add_mul _ _ _) } } end @@ -216,6 +221,15 @@ with_top.top_mul h @[simp] lemma bot_mul_bot : (⊥ * ⊥ : with_bot α) = ⊥ := with_top.top_mul_top +theorem mul_eq_bot_iff {a b : with_bot α} : a * b = ⊥ ↔ a ≠ 0 ∧ b = ⊥ ∨ a = ⊥ ∧ b ≠ 0 := +with_top.mul_eq_top_iff + +theorem bot_lt_mul' [has_lt α] {a b : with_bot α} (ha : ⊥ < a) (hb : ⊥ < b) : ⊥ < a * b := +@with_top.mul_lt_top' αᵒᵈ _ _ _ _ _ _ ha hb + +theorem bot_lt_mul [has_lt α] {a b : with_bot α} (ha : a ≠ ⊥) (hb : b ≠ ⊥) : ⊥ < a * b := +@with_top.mul_lt_top αᵒᵈ _ _ _ _ _ _ ha hb + end has_mul section mul_zero_class @@ -228,16 +242,6 @@ with_top.coe_mul lemma mul_coe {b : α} (hb : b ≠ 0) {a : with_bot α} : a * b = a.bind (λa:α, ↑(a * b)) := with_top.mul_coe hb -@[simp] lemma mul_eq_bot_iff {a b : with_bot α} : a * b = ⊥ ↔ (a ≠ 0 ∧ b = ⊥) ∨ (a = ⊥ ∧ b ≠ 0) := -with_top.mul_eq_top_iff - -lemma bot_lt_mul [preorder α] {a b : with_bot α} (ha : ⊥ < a) (hb : ⊥ < b) : ⊥ < a * b := -begin - lift a to α using ne_bot_of_gt ha, - lift b to α using ne_bot_of_gt hb, - simp only [← coe_mul, bot_lt_coe], -end - end mul_zero_class /-- `nontrivial α` is needed here as otherwise we have `1 * ⊥ = ⊥` but also `= 0 * ⊥ = 0`. -/ diff --git a/src/algebra/order/sub/with_top.lean b/src/algebra/order/sub/with_top.lean index 5e755488d2801..879e38bbf9fcc 100644 --- a/src/algebra/order/sub/with_top.lean +++ b/src/algebra/order/sub/with_top.lean @@ -33,6 +33,11 @@ instance : has_sub (with_top α) := @[simp] lemma top_sub_coe {a : α} : (⊤ : with_top α) - a = ⊤ := rfl @[simp] lemma sub_top {a : with_top α} : a - ⊤ = 0 := by { cases a; refl } +@[simp] theorem sub_eq_top_iff : ∀ {a b : with_top α}, a - b = ⊤ ↔ a = ⊤ ∧ b ≠ ⊤ +| _ ⊤ := by simp +| ⊤ (b : α) := by simp +| (a : α) (b : α) := by simp only [← coe_sub, coe_ne_top, false_and] + lemma map_sub [has_sub β] [has_zero β] {f : α → β} (h : ∀ x y, f (x - y) = f x - f y) (h₀ : f 0 = 0) : ∀ x y : with_top α, (x - y).map f = x.map f - y.map f diff --git a/src/analysis/bounded_variation.lean b/src/analysis/bounded_variation.lean index d1db929d718b9..b224ea147a178 100644 --- a/src/analysis/bounded_variation.lean +++ b/src/analysis/bounded_variation.lean @@ -933,12 +933,7 @@ lemma lipschitz_on_with.comp_has_bounded_variation_on {f : E → F} {C : ℝ≥0 (hf : lipschitz_on_with C f t) {g : α → E} {s : set α} (hg : maps_to g s t) (h : has_bounded_variation_on g s) : has_bounded_variation_on (f ∘ g) s := -begin - dsimp [has_bounded_variation_on] at h, - apply ne_of_lt, - apply (hf.comp_evariation_on_le hg).trans_lt, - simp [lt_top_iff_ne_top, h], -end +ne_top_of_le_ne_top (ennreal.mul_ne_top ennreal.coe_ne_top h) (hf.comp_evariation_on_le hg) lemma lipschitz_on_with.comp_has_locally_bounded_variation_on {f : E → F} {C : ℝ≥0} {t : set E} (hf : lipschitz_on_with C f t) {g : α → E} {s : set α} (hg : maps_to g s t) diff --git a/src/analysis/mean_inequalities_pow.lean b/src/analysis/mean_inequalities_pow.lean index 0865f42f60bc1..57f964e83f40d 100644 --- a/src/analysis/mean_inequalities_pow.lean +++ b/src/analysis/mean_inequalities_pow.lean @@ -198,7 +198,7 @@ begin have hp_not_nonpos : ¬ p ≤ 0, by simp [hp_pos], have hp_not_neg : ¬ p < 0, by simp [hp_nonneg], have h_top_iff_rpow_top : ∀ (i : ι) (hi : i ∈ s), w i * z i = ⊤ ↔ w i * (z i) ^ p = ⊤, - by simp [hp_pos, hp_nonneg, hp_not_nonpos, hp_not_neg], + by simp [ennreal.mul_eq_top, hp_pos, hp_nonneg, hp_not_nonpos, hp_not_neg], refine le_of_top_imp_top_of_to_nnreal_le _ _, { -- first, prove `(∑ i in s, w i * z i) ^ p = ⊤ → ∑ i in s, (w i * z i ^ p) = ⊤` rw [rpow_eq_top_iff, sum_eq_top_iff, sum_eq_top_iff], diff --git a/src/analysis/p_series.lean b/src/analysis/p_series.lean index d6b509255ec30..bc31520561bd5 100644 --- a/src/analysis/p_series.lean +++ b/src/analysis/p_series.lean @@ -122,7 +122,7 @@ begin split; intro h, { replace hf : ∀ m n, 1 < m → m ≤ n → (f n : ℝ≥0∞) ≤ f m := λ m n hm hmn, ennreal.coe_le_coe.2 (hf (zero_lt_one.trans hm) hmn), - simpa [h, ennreal.add_eq_top] using (ennreal.tsum_condensed_le hf) }, + simpa [h, ennreal.add_eq_top, ennreal.mul_eq_top] using ennreal.tsum_condensed_le hf }, { replace hf : ∀ m n, 0 < m → m ≤ n → (f n : ℝ≥0∞) ≤ f m := λ m n hm hmn, ennreal.coe_le_coe.2 (hf hm hmn), simpa [h, ennreal.add_eq_top] using (ennreal.le_tsum_condensed hf) } diff --git a/src/data/real/ennreal.lean b/src/data/real/ennreal.lean index d7b87d244de5d..c35139f2da10b 100644 --- a/src/data/real/ennreal.lean +++ b/src/data/real/ennreal.lean @@ -720,28 +720,30 @@ mul_right_mono.map_max lemma mul_max : a * max b c = max (a * b) (a * c) := mul_left_mono.map_max -lemma mul_eq_mul_left : a ≠ 0 → a ≠ ∞ → (a * b = a * c ↔ b = c) := +theorem mul_left_strictMono (h0 : a ≠ 0) (hinf : a ≠ ∞) : strict_mono ((*) a) := begin - cases a; cases b; cases c; - simp [none_eq_top, some_eq_coe, mul_top, top_mul, -coe_mul, coe_mul.symm, - nnreal.mul_eq_mul_left] {contextual := tt}, + lift a to ℝ≥0 using hinf, + rw [coe_ne_zero] at h0, + intros x y h, + contrapose! h, + simpa only [← mul_assoc, ← coe_mul, inv_mul_cancel h0, coe_one, one_mul] + using mul_le_mul' (le_refl ↑a⁻¹) h end +lemma mul_eq_mul_left (h₀ : a ≠ 0) (hinf : a ≠ ∞) : a * b = a * c ↔ b = c := +(mul_left_strictMono h₀ hinf).injective.eq_iff + lemma mul_eq_mul_right : c ≠ 0 → c ≠ ∞ → (a * c = b * c ↔ a = b) := mul_comm c a ▸ mul_comm c b ▸ mul_eq_mul_left -lemma mul_le_mul_left : a ≠ 0 → a ≠ ∞ → (a * b ≤ a * c ↔ b ≤ c) := -begin - cases a; cases b; cases c; - simp [none_eq_top, some_eq_coe, mul_top, top_mul, -coe_mul, coe_mul.symm] {contextual := tt}, - assume h, exact mul_le_mul_left (pos_iff_ne_zero.2 h) -end +lemma mul_le_mul_left (h₀ : a ≠ 0) (hinf : a ≠ ∞) : a * b ≤ a * c ↔ b ≤ c := +(mul_left_strictMono h₀ hinf).le_iff_le lemma mul_le_mul_right : c ≠ 0 → c ≠ ∞ → (a * c ≤ b * c ↔ a ≤ b) := mul_comm c a ▸ mul_comm c b ▸ mul_le_mul_left -lemma mul_lt_mul_left : a ≠ 0 → a ≠ ∞ → (a * b < a * c ↔ b < c) := -λ h0 ht, by simp only [mul_le_mul_left h0 ht, lt_iff_le_not_le] +lemma mul_lt_mul_left (h₀ : a ≠ 0) (hinf : a ≠ ∞) : a * b < a * c ↔ b < c := +(mul_left_strictMono h₀ hinf).lt_iff_lt lemma mul_lt_mul_right : c ≠ 0 → c ≠ ∞ → (a * c < b * c ↔ a < b) := mul_comm c a ▸ mul_comm c b ▸ mul_lt_mul_left @@ -749,6 +751,7 @@ mul_comm c a ▸ mul_comm c b ▸ mul_lt_mul_left end mul section cancel + /-- An element `a` is `add_le_cancellable` if `a + b ≤ a + c` implies `b ≤ c` for all `b` and `c`. This is true in `ℝ≥0∞` for all elements except `∞`. -/ lemma add_le_cancellable_iff_ne {a : ℝ≥0∞} : add_le_cancellable a ↔ a ≠ ∞ := diff --git a/src/data/real/ereal.lean b/src/data/real/ereal.lean index 7007030abf71c..9cf9844ba58d6 100644 --- a/src/data/real/ereal.lean +++ b/src/data/real/ereal.lean @@ -405,7 +405,8 @@ by cases x; cases y; refl exact nnreal.coe_pos.2 (bot_lt_iff_ne_bot.2 h) }, simp only [coe_nnreal_eq_coe_real, coe_ennreal_top, (*), ereal.mul, A, if_true] } end -| (x : ℝ≥0) (y : ℝ≥0) := by simp [←ennreal.coe_mul, coe_nnreal_eq_coe_real] +| (x : ℝ≥0) (y : ℝ≥0) := by simp only [← ennreal.coe_mul, coe_nnreal_eq_coe_real, + nnreal.coe_mul, ereal.coe_mul] @[norm_cast] lemma coe_ennreal_nsmul (n : ℕ) (x : ℝ≥0∞) : (↑(n • x) : ereal) = n • x := map_nsmul (⟨coe, coe_ennreal_zero, coe_ennreal_add⟩ : ℝ≥0∞ →+ ereal) _ _ diff --git a/src/measure_theory/function/l1_space.lean b/src/measure_theory/function/l1_space.lean index b63912f5a5cf6..379785338a518 100644 --- a/src/measure_theory/function/l1_space.lean +++ b/src/measure_theory/function/l1_space.lean @@ -147,7 +147,8 @@ has_finite_integral_congr' $ h.fun_comp norm lemma has_finite_integral_const_iff {c : β} : has_finite_integral (λ x : α, c) μ ↔ c = 0 ∨ μ univ < ∞ := -by simp [has_finite_integral, lintegral_const, lt_top_iff_ne_top, or_iff_not_imp_left] +by simp [has_finite_integral, lintegral_const, lt_top_iff_ne_top, ennreal.mul_eq_top, + or_iff_not_imp_left] lemma has_finite_integral_const [is_finite_measure μ] (c : β) : has_finite_integral (λ x : α, c) μ := diff --git a/src/measure_theory/function/lp_space.lean b/src/measure_theory/function/lp_space.lean index f803c1a4891fd..3f43ae6135b10 100644 --- a/src/measure_theory/function/lp_space.lean +++ b/src/measure_theory/function/lp_space.lean @@ -1375,11 +1375,8 @@ end lemma mem_ℒp.of_le_mul {f : α → E} {g : α → F} {c : ℝ} (hg : mem_ℒp g p μ) (hf : ae_strongly_measurable f μ) (hfg : ∀ᵐ x ∂μ, ‖f x‖ ≤ c * ‖g x‖) : mem_ℒp f p μ := -begin - simp only [mem_ℒp, hf, true_and], - apply lt_of_le_of_lt (snorm_le_mul_snorm_of_ae_le_mul hfg p), - simp [lt_top_iff_ne_top, hg.snorm_ne_top], -end +⟨hf, lt_of_le_of_lt (snorm_le_mul_snorm_of_ae_le_mul hfg p) $ + ennreal.mul_lt_top ennreal.of_real_ne_top hg.snorm_ne_top⟩ end monotonicity @@ -1682,17 +1679,13 @@ by rw [norm_def, norm_def, snorm_congr_ae (coe_fn_neg _), snorm_neg] lemma norm_le_mul_norm_of_ae_le_mul {c : ℝ} {f : Lp E p μ} {g : Lp F p μ} (h : ∀ᵐ x ∂μ, ‖f x‖ ≤ c * ‖g x‖) : ‖f‖ ≤ c * ‖g‖ := begin - by_cases pzero : p = 0, - { simp [pzero, norm_def] }, + simp only [norm_def], cases le_or_lt 0 c with hc hc, - { have := snorm_le_mul_snorm_aux_of_nonneg h hc p, - rw [← ennreal.to_real_le_to_real, ennreal.to_real_mul, ennreal.to_real_of_real hc] at this, - { exact this }, + { have := snorm_le_mul_snorm_of_ae_le_mul h p, + rwa [← ennreal.to_real_le_to_real, ennreal.to_real_mul, ennreal.to_real_of_real hc] at this, { exact (Lp.mem_ℒp _).snorm_ne_top }, - { simp [(Lp.mem_ℒp _).snorm_ne_top] } }, + { exact ennreal.mul_ne_top ennreal.of_real_ne_top (Lp.mem_ℒp _).snorm_ne_top } }, { have := snorm_le_mul_snorm_aux_of_neg h hc p, - simp only [snorm_eq_zero_iff (Lp.ae_strongly_measurable _) pzero, ← eq_zero_iff_ae_eq_zero] - at this, simp [this] } end diff --git a/src/measure_theory/integral/mean_inequalities.lean b/src/measure_theory/integral/mean_inequalities.lean index e323bac178bdf..1e1c756200284 100644 --- a/src/measure_theory/integral/mean_inequalities.lean +++ b/src/measure_theory/integral/mean_inequalities.lean @@ -114,7 +114,8 @@ begin end ... ≤ npf * nqg : begin - rw lintegral_mul_const' (npf * nqg) _ (by simp [hf_nontop, hg_nontop, hf_nonzero, hg_nonzero]), + rw lintegral_mul_const' (npf * nqg) _ + (by simp [hf_nontop, hg_nontop, hf_nonzero, hg_nonzero, ennreal.mul_eq_top]), nth_rewrite 1 ←one_mul (npf * nqg), refine mul_le_mul _ (le_refl (npf * nqg)), have hf1 := lintegral_rpow_fun_mul_inv_snorm_eq_one hpq.pos hf_nonzero hf_nontop, diff --git a/src/order/with_bot.lean b/src/order/with_bot.lean index b0cefb0049796..2ad1002fe61ca 100644 --- a/src/order/with_bot.lean +++ b/src/order/with_bot.lean @@ -82,6 +82,17 @@ def unbot' (d : α) (x : with_bot α) : α := rec_bot_coe d id x @[norm_cast] lemma coe_eq_coe : (a : with_bot α) = b ↔ a = b := option.some_inj +lemma unbot'_eq_iff {d y : α} {x : with_bot α} : unbot' d x = y ↔ x = y ∨ x = ⊥ ∧ y = d := +by induction x using with_bot.rec_bot_coe; simp [@eq_comm _ d, coe_eq_coe] + +@[simp] +theorem unbot'_eq_self_iff {d : α} {x : with_bot α} : unbot' d x = d ↔ x = d ∨ x = ⊥ := +by simp [unbot'_eq_iff] + +theorem unbot'_eq_unbot'_iff {d : α} {x y : with_bot α} : + unbot' d x = unbot' d y ↔ x = y ∨ x = d ∧ y = ⊥ ∨ x = ⊥ ∧ y = d := +by induction y using with_bot.rec_bot_coe; simp [unbot'_eq_iff, or_comm, coe_eq_coe] + /-- Lift a map `f : α → β` to `with_bot α → with_bot β`. Implemented using `option.map`. -/ def map (f : α → β) : with_bot α → with_bot β := option.map f @@ -172,6 +183,12 @@ lemma lt_coe_iff : ∀ {x : with_bot α}, x < b ↔ ∀ a, x = ↑a → a < b | (some b) := by simp [some_eq_coe, coe_eq_coe, coe_lt_coe] | none := by simp [none_eq_bot, bot_lt_coe] +/-- A version of `bot_lt_iff_ne_bot` for `with_bot` that only requires `has_lt α`, not +`partial_order α`. -/ +protected theorem bot_lt_iff_ne_bot : ∀ {x : with_bot α}, ⊥ < x ↔ x ≠ ⊥ +| ⊥ := iff_of_false (with_bot.not_lt_none _) (λ h, h rfl) +| (x : α) := by simp [bot_lt_coe] + end has_lt instance [preorder α] : preorder (with_bot α) := @@ -462,6 +479,16 @@ def untop' (d : α) (x : with_top α) : α := rec_top_coe d id x @[norm_cast] lemma coe_eq_coe : (a : with_top α) = b ↔ a = b := option.some_inj +lemma untop'_eq_iff {d y : α} {x : with_top α} : untop' d x = y ↔ x = y ∨ x = ⊤ ∧ y = d := +with_bot.unbot'_eq_iff + +@[simp] theorem untop'_eq_self_iff {d : α} {x : with_top α} : untop' d x = d ↔ x = d ∨ x = ⊤ := +with_bot.unbot'_eq_self_iff + +theorem untop'_eq_untop'_iff {d : α} {x y : with_top α} : + untop' d x = untop' d y ↔ x = y ∨ x = d ∧ y = ⊤ ∨ x = ⊤ ∧ y = d := +with_bot.unbot'_eq_unbot'_iff + /-- Lift a map `f : α → β` to `with_top α → with_top β`. Implemented using `option.map`. -/ def map (f : α → β) : with_top α → with_top β := option.map f @@ -679,6 +706,11 @@ begin exact forall₂_congr (λ _ _, iff.rfl) end +/-- A version of `lt_top_iff_ne_top` for `with_top` that only requires `has_lt α`, not +`partial_order α`. -/ +protected theorem lt_top_iff_ne_top {x : with_top α} : x < ⊤ ↔ x ≠ ⊤ := +@with_bot.bot_lt_iff_ne_bot αᵒᵈ _ x + end has_lt instance [preorder α] : preorder (with_top α) := diff --git a/src/topology/instances/ennreal.lean b/src/topology/instances/ennreal.lean index 981d10c97a9dc..49c3479737fbd 100644 --- a/src/topology/instances/ennreal.lean +++ b/src/topology/instances/ennreal.lean @@ -1424,7 +1424,7 @@ begin end⟩ lemma continuous_of_le_add_edist {f : α → ℝ≥0∞} (C : ℝ≥0∞) - (hC : C ≠ ⊤) (h : ∀x y, f x ≤ f y + C * edist x y) : continuous f := + (hC : C ≠ ⊤) (h : ∀ x y, f x ≤ f y + C * edist x y) : continuous f := begin rcases eq_or_ne C 0 with (rfl|C0), { simp only [zero_mul, add_zero] at h, @@ -1434,7 +1434,7 @@ begin { have : f =ᶠ[𝓝 x] (λ _, ∞), { filter_upwards [emetric.ball_mem_nhds x ennreal.coe_lt_top], refine λ y (hy : edist y x < ⊤), _, rw edist_comm at hy, - simpa [hx, hC, hy.ne] using h x y }, + simpa [hx, ennreal.mul_ne_top hC hy.ne] using h x y }, exact this.continuous_at }, { refine (ennreal.tendsto_nhds hx).2 (λ ε (ε0 : 0 < ε), _), filter_upwards [emetric.closed_ball_mem_nhds x (ennreal.div_pos_iff.2 ⟨ε0.ne', hC⟩)], diff --git a/src/topology/metric_space/hausdorff_distance.lean b/src/topology/metric_space/hausdorff_distance.lean index 91616d78d4785..2a93ac1486ed0 100644 --- a/src/topology/metric_space/hausdorff_distance.lean +++ b/src/topology/metric_space/hausdorff_distance.lean @@ -1123,9 +1123,10 @@ lemma diam_cthickening_le {α : Type*} [pseudo_metric_space α] (s : set α) (h begin by_cases hs : bounded (cthickening ε s), { replace hs := hs.mono (self_subset_cthickening _), - have : (2 : ℝ≥0∞) * @coe ℝ≥0 _ _ ⟨ε, hε⟩ ≠ ⊤ := by simp, + lift ε to ℝ≥0 using hε, + have : (2 : ℝ≥0∞) * ε ≠ ⊤ := by simp [ennreal.mul_eq_top], refine (ennreal.to_real_mono (ennreal.add_ne_top.2 ⟨hs.ediam_ne_top, this⟩) $ - ediam_cthickening_le ⟨ε, hε⟩).trans_eq _, + ediam_cthickening_le ε).trans_eq _, simp [ennreal.to_real_add hs.ediam_ne_top this, diam] }, { rw diam_eq_zero_of_unbounded hs, positivity } From 75608affb24b4f48699fbcd38f227827f7793771 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 24 Feb 2023 15:00:45 +0000 Subject: [PATCH 0525/1291] chore(data/set/basic): Golf #18356 (#18491) and fix a docstring that was made a comment --- src/data/set/basic.lean | 19 +++---------------- 1 file changed, 3 insertions(+), 16 deletions(-) diff --git a/src/data/set/basic.lean b/src/data/set/basic.lean index 729ea6cb1e1db..279f8b6e1c8b3 100644 --- a/src/data/set/basic.lean +++ b/src/data/set/basic.lean @@ -698,21 +698,8 @@ begin exacts [(ha hx).elim, hxt] end -theorem subset_insert_iff_of_not_mem {s t : set α} {a : α} (h : a ∉ s) : - s ⊆ insert a t ↔ s ⊆ t := -begin - constructor, - { intros g y hy, - specialize g hy, - rw [mem_insert_iff] at g, - rcases g with g | g, - { rw [g] at hy, - contradiction }, - { assumption }}, - { intros g y hy, - specialize g hy, - exact mem_insert_of_mem _ g } -end +theorem subset_insert_iff_of_not_mem (ha : a ∉ s) : s ⊆ insert a t ↔ s ⊆ t := +forall₂_congr $ λ b hb, or_iff_right $ ne_of_mem_of_not_mem hb ha theorem ssubset_iff_insert {s t : set α} : s ⊂ t ↔ ∃ a ∉ s, insert a s ⊆ t := begin @@ -1317,7 +1304,7 @@ ext $ λ s, subset_empty_iff @[simp] theorem powerset_univ : 𝒫 (univ : set α) = univ := eq_univ_of_forall subset_univ -/- The powerset of a singleton contains only `∅` and the singleton itself. -/ +/-- The powerset of a singleton contains only `∅` and the singleton itself. -/ theorem powerset_singleton (x : α) : 𝒫 ({x} : set α) = {∅, {x}} := by { ext y, rw [mem_powerset_iff, subset_singleton_iff_eq, mem_insert_iff, mem_singleton_iff] } From 858a10cf68fd6c06872950fc58c4dcc68d465591 Mon Sep 17 00:00:00 2001 From: datokrat Date: Fri, 24 Feb 2023 18:18:47 +0000 Subject: [PATCH 0526/1291] feat(analysis/convex/body): endow with Hausdorff metric (#16338) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Endows the type of convex bodies with the Hausdorff pseudo-metric. If the underlying vector space is normed instead of only seminormed, this gives rise to the Hausdorff metric. Co-authored-by: Yaël Dillies --- src/analysis/convex/body.lean | 77 ++++++++++++++++++++++++++++------- 1 file changed, 62 insertions(+), 15 deletions(-) diff --git a/src/analysis/convex/body.lean b/src/analysis/convex/body.lean index ec6b5060befa9..8a8698e4986b1 100644 --- a/src/analysis/convex/body.lean +++ b/src/analysis/convex/body.lean @@ -4,23 +4,25 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Paul A. Reichert -/ import analysis.convex.basic -import data.real.nnreal -import topology.algebra.module.basic -import topology.instances.real +import analysis.normed_space.basic +import topology.metric_space.hausdorff_distance /-! -# convex bodies +# Convex bodies This file contains the definition of the type `convex_body V` consisting of convex, compact, nonempty subsets of a real topological vector space `V`. -`convex_body V` is a module over the nonnegative reals (`nnreal`). +`convex_body V` is a module over the nonnegative reals (`nnreal`) and a pseudo-metric space. +If `V` is a normed space, `convex_body V` is a metric space. + +## TODO -TODOs: -- endow it with the Hausdorff metric - define positive convex bodies, requiring the interior to be nonempty - introduce support sets +- Characterise the interaction of the distance with algebraic operations, eg + `dist (a • K) (a • L) = ‖a‖ * dist K L`, `dist (a +ᵥ K) (a +ᵥ L) = dist K L` ## Tags @@ -30,30 +32,29 @@ convex, convex body open_locale pointwise open_locale nnreal -variables (V : Type*) [topological_space V] [add_comm_group V] [has_continuous_add V] - [module ℝ V] [has_continuous_smul ℝ V] +variables {V : Type*} /-- Let `V` be a real topological vector space. A subset of `V` is a convex body if and only if it is convex, compact, and nonempty. -/ -structure convex_body := +structure convex_body (V : Type*) [topological_space V] [add_comm_monoid V] [has_smul ℝ V] := (carrier : set V) (convex' : convex ℝ carrier) (is_compact' : is_compact carrier) (nonempty' : carrier.nonempty) namespace convex_body - -variables {V} +section TVS +variables [topological_space V] [add_comm_group V] [module ℝ V] instance : set_like (convex_body V) V := { coe := convex_body.carrier, coe_injective' := λ K L h, by { cases K, cases L, congr' } } -lemma convex (K : convex_body V) : convex ℝ (K : set V) := K.convex' -lemma is_compact (K : convex_body V) : is_compact (K : set V) := K.is_compact' -lemma nonempty (K : convex_body V) : (K : set V).nonempty := K.nonempty' +protected lemma convex (K : convex_body V) : convex ℝ (K : set V) := K.convex' +protected lemma is_compact (K : convex_body V) : is_compact (K : set V) := K.is_compact' +protected lemma nonempty (K : convex_body V) : (K : set V).nonempty := K.nonempty' @[ext] protected lemma ext {K L : convex_body V} (h : (K : set V) = L) : K = L := set_like.ext' h @@ -61,6 +62,9 @@ protected lemma ext {K L : convex_body V} (h : (K : set V) = L) : K = L := set_l @[simp] lemma coe_mk (s : set V) (h₁ h₂ h₃) : (mk s h₁ h₂ h₃ : set V) = s := rfl +section has_continuous_add +variables [has_continuous_add V] + instance : add_monoid (convex_body V) := -- we cannot write K + L to avoid reducibility issues with the set.has_add instance { add := λ K L, ⟨set.image2 (+) K L, @@ -84,12 +88,18 @@ instance : add_comm_monoid (convex_body V) := { add_comm := λ K L, by { ext, simp only [coe_add, add_comm] }, .. convex_body.add_monoid } +end has_continuous_add + +variables [has_continuous_smul ℝ V] + instance : has_smul ℝ (convex_body V) := { smul := λ c K, ⟨c • (K : set V), K.convex.smul _, K.is_compact.smul _, K.nonempty.smul_set⟩ } @[simp] lemma coe_smul (c : ℝ) (K : convex_body V) : (↑(c • K) : set V) = c • (K : set V) := rfl +variables [has_continuous_add V] + instance : distrib_mul_action ℝ (convex_body V) := { to_has_smul := convex_body.has_smul, one_smul := λ K, by { ext, simp only [coe_smul, one_smul] }, @@ -112,4 +122,41 @@ instance : module ℝ≥0 (convex_body V) := end, zero_smul := λ K, by { ext1, exact set.zero_smul_set K.nonempty } } +end TVS + +section seminormed_add_comm_group +variables [seminormed_add_comm_group V] [normed_space ℝ V] (K L : convex_body V) + +protected lemma bounded : metric.bounded (K : set V) := K.is_compact.bounded + +lemma Hausdorff_edist_ne_top {K L : convex_body V} : emetric.Hausdorff_edist (K : set V) L ≠ ⊤ := +by apply_rules [metric.Hausdorff_edist_ne_top_of_nonempty_of_bounded, convex_body.nonempty, + convex_body.bounded] + +/-- Convex bodies in a fixed seminormed space $V$ form a pseudo-metric space under the Hausdorff +metric. -/ +noncomputable instance : pseudo_metric_space (convex_body V) := +{ dist := λ K L, metric.Hausdorff_dist (K : set V) L, + dist_self := λ _, metric.Hausdorff_dist_self_zero, + dist_comm := λ _ _, metric.Hausdorff_dist_comm, + dist_triangle := λ K L M, metric.Hausdorff_dist_triangle Hausdorff_edist_ne_top } + +@[simp, norm_cast] +lemma Hausdorff_dist_coe : metric.Hausdorff_dist (K : set V) L = dist K L := rfl + +@[simp, norm_cast] lemma Hausdorff_edist_coe : emetric.Hausdorff_edist (K : set V) L = edist K L := +by { rw edist_dist, exact (ennreal.of_real_to_real Hausdorff_edist_ne_top).symm } + +end seminormed_add_comm_group + +section normed_add_comm_group +variables [normed_add_comm_group V] [normed_space ℝ V] + +/-- Convex bodies in a fixed normed space `V` form a metric space under the Hausdorff metric. -/ +noncomputable instance : metric_space (convex_body V) := +{ eq_of_dist_eq_zero := λ K L hd, convex_body.ext $ + (K.is_compact.is_closed.Hausdorff_dist_zero_iff_eq + L.is_compact.is_closed Hausdorff_edist_ne_top).mp hd } + +end normed_add_comm_group end convex_body From 733fa0048f88bd38678c283c8c1bb1445ac5e23b Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Fri, 24 Feb 2023 18:18:48 +0000 Subject: [PATCH 0527/1291] feat(measure_theory/integral): integration lemmas from #18392 (#18466) Some lemmas about integration of continuous functions split off from my Poisson-summation / zeta-functions project at #18392. --- .../integral/interval_integral.lean | 24 ++++++++++++++ src/measure_theory/integral/set_integral.lean | 32 +++++++++++++++++-- src/measure_theory/measure/lebesgue.lean | 29 +++++++++++++++++ 3 files changed, 83 insertions(+), 2 deletions(-) diff --git a/src/measure_theory/integral/interval_integral.lean b/src/measure_theory/integral/interval_integral.lean index 2747eebf1519c..81dd304d6c1b2 100644 --- a/src/measure_theory/integral/interval_integral.lean +++ b/src/measure_theory/integral/interval_integral.lean @@ -1065,6 +1065,30 @@ begin end open topological_space + +/-- Interval integrals commute with countable sums, when the supremum norms are summable (a +special case of the dominated convergence theorem). -/ +lemma has_sum_interval_integral_of_summable_norm [countable ι] {f : ι → C(ℝ, E)} + (hf_sum : summable (λ i : ι, ‖(f i).restrict (⟨uIcc a b, is_compact_uIcc⟩ : compacts ℝ)‖)) : + has_sum (λ i : ι, ∫ x in a..b, f i x) (∫ x in a..b, (∑' i : ι, f i x)) := +begin + refine has_sum_integral_of_dominated_convergence + (λ i (x : ℝ), ‖(f i).restrict ↑(⟨uIcc a b, is_compact_uIcc⟩ : compacts ℝ)‖) + (λ i, (map_continuous $ f i).ae_strongly_measurable) + (λ i, ae_of_all _ (λ x hx, ((f i).restrict ↑(⟨uIcc a b, is_compact_uIcc⟩ : + compacts ℝ)).norm_coe_le_norm ⟨x, ⟨hx.1.le, hx.2⟩⟩)) + (ae_of_all _ (λ x hx, hf_sum)) + interval_integrable_const + (ae_of_all _ (λ x hx, summable.has_sum _)), + -- next line is slow, & doesn't work with "exact" in place of "apply" -- ? + apply continuous_map.summable_apply (summable_of_summable_norm hf_sum) ⟨x, ⟨hx.1.le, hx.2⟩⟩, +end + +lemma tsum_interval_integral_eq_of_summable_norm [countable ι] {f : ι → C(ℝ, E)} + (hf_sum : summable (λ i : ι, ‖(f i).restrict (⟨uIcc a b, is_compact_uIcc⟩ : compacts ℝ)‖)) : + ∑' (i : ι), ∫ x in a..b, f i x = ∫ x in a..b, (∑' i : ι, f i x) := +(has_sum_interval_integral_of_summable_norm hf_sum).tsum_eq + variables {X : Type*} [topological_space X] [first_countable_topology X] /-- Continuity of interval integral with respect to a parameter, at a point within a set. diff --git a/src/measure_theory/integral/set_integral.lean b/src/measure_theory/integral/set_integral.lean index 0432a39540c6a..37a87b63bcd4d 100644 --- a/src/measure_theory/integral/set_integral.lean +++ b/src/measure_theory/integral/set_integral.lean @@ -5,8 +5,10 @@ Authors: Zhouhang Zhou, Yury Kudryashov -/ import measure_theory.integral.integrable_on import measure_theory.integral.bochner +import measure_theory.function.locally_integrable import order.filter.indicator_function import topology.metric_space.thickened_indicator +import topology.continuous_function.compact /-! # Set integral @@ -711,9 +713,9 @@ end nonneg section integrable_Union -variables {μ : measure α} [normed_add_comm_group E] {f : α → E} [countable β] {s : β → set α} +variables {μ : measure α} [normed_add_comm_group E] [countable β] -lemma integrable_on_Union_of_summable_integral_norm +lemma integrable_on_Union_of_summable_integral_norm {f : α → E} {s : β → set α} (hs : ∀ (b : β), measurable_set (s b)) (hi : ∀ (b : β), integrable_on f (s b) μ) (h : summable (λ (b : β), ∫ (a : α) in s b, ‖f a‖ ∂μ)) : integrable_on f (Union s) μ := @@ -729,6 +731,32 @@ begin convert ennreal.of_real_lt_top, end +variables [topological_space α] [borel_space α] [metrizable_space α] [is_locally_finite_measure μ] + +/-- If `s` is a countable family of compact sets, `f` is a continuous function, and the sequence +`‖f.restrict (s i)‖ * μ (s i)` is summable, then `f` is integrable on the union of the `s i`. -/ +lemma integrable_on_Union_of_summable_norm_restrict {f : C(α, E)} {s : β → compacts α} + (hf : summable (λ i : β, ‖f.restrict (s i)‖ * ennreal.to_real (μ $ s i))) : + integrable_on f (⋃ i : β, s i) μ := +begin + refine integrable_on_Union_of_summable_integral_norm + (λ i, (s i).is_compact.is_closed.measurable_set) + (λ i, (map_continuous f).continuous_on.integrable_on_compact (s i).is_compact) + (summable_of_nonneg_of_le (λ ι, integral_nonneg (λ x, norm_nonneg _)) (λ i, _) hf), + rw ←(real.norm_of_nonneg (integral_nonneg (λ a, norm_nonneg _)) : ‖_‖ = ∫ x in s i, ‖f x‖ ∂μ), + exact norm_set_integral_le_of_norm_le_const' (s i).is_compact.measure_lt_top + (s i).is_compact.is_closed.measurable_set + (λ x hx, (norm_norm (f x)).symm ▸ (f.restrict ↑(s i)).norm_coe_le_norm ⟨x, hx⟩) +end + +/-- If `s` is a countable family of compact sets covering `α`, `f` is a continuous function, and +the sequence `‖f.restrict (s i)‖ * μ (s i)` is summable, then `f` is integrable. -/ +lemma integrable_of_summable_norm_restrict {f : C(α, E)} {s : β → compacts α} + (hf : summable (λ i : β, ‖f.restrict (s i)‖ * ennreal.to_real (μ $ s i))) + (hs : (⋃ i : β, ↑(s i)) = (univ : set α)) : + integrable f μ := +by simpa only [hs, integrable_on_univ] using integrable_on_Union_of_summable_norm_restrict hf + end integrable_Union section tendsto_mono diff --git a/src/measure_theory/measure/lebesgue.lean b/src/measure_theory/measure/lebesgue.lean index 2c8112475e211..fde7b7ecdfdbe 100644 --- a/src/measure_theory/measure/lebesgue.lean +++ b/src/measure_theory/measure/lebesgue.lean @@ -601,3 +601,32 @@ begin apply h'x p pA ⟨xs, xp⟩ }, { exact false.elim (hx ⟨xs, Hx⟩) } end + +section summable_norm_Icc + +open continuous_map + +/- The following lemma is a minor variation on `integrable_of_summable_norm_restrict` in +`measure_theory.integral.set_integral`, but it is placed here because it needs to know that +`Icc a b` has volume `b - a`. -/ + +/-- If the sequence with `n`-th term the the sup norm of `λ x, f (x + n)` on the interval `Icc 0 1`, +for `n ∈ ℤ`, is summable, then `f` is integrable on `ℝ`. -/ +lemma real.integrable_of_summable_norm_Icc {E : Type*} [normed_add_comm_group E] {f : C(ℝ, E)} + (hf : summable (λ n : ℤ, ‖(f.comp $ continuous_map.add_right n).restrict (Icc 0 1)‖)) : + integrable f := +begin + refine integrable_of_summable_norm_restrict (summable_of_nonneg_of_le + (λ n : ℤ, mul_nonneg (norm_nonneg (f.restrict (⟨Icc n (n + 1), is_compact_Icc⟩ : compacts ℝ))) + ennreal.to_real_nonneg) (λ n, _) hf) (Union_Icc_int_cast ℝ), + simp only [compacts.coe_mk, real.volume_Icc, add_sub_cancel', ennreal.to_real_of_real zero_le_one, + mul_one, norm_le _ (norm_nonneg _)], + intro x, + have := ((f.comp $ continuous_map.add_right n).restrict (Icc 0 1)).norm_coe_le_norm + ⟨x - n, ⟨sub_nonneg.mpr x.2.1, sub_le_iff_le_add'.mpr x.2.2⟩⟩, + simpa only [continuous_map.restrict_apply, comp_apply, coe_add_right, subtype.coe_mk, + sub_add_cancel] + using this, +end + +end summable_norm_Icc From 271bf175e6c51b8d31d6c0107b7bb4a967c7277e Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Fri, 24 Feb 2023 18:18:51 +0000 Subject: [PATCH 0528/1291] feat(number_theory/number_field/embeddings): some additional lemmas (#18473) Add some small lemmas that will be useful for other PRs. --- .../number_field/embeddings.lean | 43 ++++++++++++++++--- 1 file changed, 37 insertions(+), 6 deletions(-) diff --git a/src/number_theory/number_field/embeddings.lean b/src/number_theory/number_field/embeddings.lean index e5446b71d7756..cf7b9af532659 100644 --- a/src/number_theory/number_field/embeddings.lean +++ b/src/number_theory/number_field/embeddings.lean @@ -10,6 +10,7 @@ import number_theory.number_field.basic import ring_theory.norm import topology.instances.complex + /-! # Embeddings of number fields This file defines the embeddings of a number field into an algebraic closed field. @@ -238,11 +239,24 @@ lemma apply (φ : K →+* ℂ) (x : K) : (mk φ) x = complex.abs (φ x) := rfl /-- For an infinite place `w`, return an embedding `φ` such that `w = infinite_place φ` . -/ noncomputable def embedding (w : infinite_place K) : K →+* ℂ := (w.2).some +@[simp] lemma mk_embedding (w : infinite_place K) : mk (embedding w) = w := subtype.ext (w.2).some_spec -lemma pos_iff (w : infinite_place K) (x : K) : 0 < w x ↔ x ≠ 0 := absolute_value.pos_iff w.1 +@[simp] +lemma abs_embedding (w : infinite_place K) (x : K) : + complex.abs (embedding w x) = w x := congr_fun (congr_arg coe_fn w.2.some_spec) x + +lemma eq_iff_eq (x : K) (r : ℝ) : + (∀ w : infinite_place K, w x = r) ↔ (∀ φ : K →+* ℂ, ‖φ x‖ = r) := +⟨λ hw φ, hw (mk φ), λ hφ ⟨w, ⟨φ, rfl⟩⟩, hφ φ⟩ + +lemma le_iff_le (x : K) (r : ℝ) : + (∀ w : infinite_place K, w x ≤ r) ↔ (∀ φ : K →+* ℂ, ‖φ x‖ ≤ r) := +⟨λ hw φ, hw (mk φ), λ hφ ⟨w, ⟨φ, rfl⟩⟩, hφ φ⟩ + +lemma pos_iff {w : infinite_place K} {x : K} : 0 < w x ↔ x ≠ 0 := absolute_value.pos_iff w.1 @[simp] lemma mk_conjugate_eq (φ : K →+* ℂ) : @@ -293,6 +307,7 @@ def is_real (w : infinite_place K) : Prop := def is_complex (w : infinite_place K) : Prop := ∃ φ : K →+* ℂ, ¬ complex_embedding.is_real φ ∧ mk φ = w +@[simp] lemma _root_.number_field.complex_embeddings.is_real.embedding_mk {φ : K →+* ℂ} (h : complex_embedding.is_real φ) : embedding (mk φ) = φ := @@ -354,13 +369,29 @@ noncomputable def mk_complex : {φ : K →+* ℂ // ¬ complex_embedding.is_real φ} → {w : infinite_place K // is_complex w} := subtype.map mk (λ φ hφ, ⟨φ, hφ, rfl⟩) +lemma mk_complex_embedding (φ : {φ : K →+* ℂ // ¬ complex_embedding.is_real φ}) : + ((mk_complex K φ) : infinite_place K).embedding = φ ∨ + ((mk_complex K φ) : infinite_place K).embedding = complex_embedding.conjugate φ := +begin + rw [@eq_comm _ _ ↑φ, @eq_comm _ _ (complex_embedding.conjugate ↑φ), ← mk_eq_iff, mk_embedding], + refl, +end + +@[simp] +lemma mk_real_coe (φ : {φ : K →+* ℂ // complex_embedding.is_real φ}) : + (mk_real K φ : infinite_place K) = mk (φ : K →+* ℂ) := rfl + +@[simp] +lemma mk_complex_coe (φ : {φ : K →+* ℂ // ¬ complex_embedding.is_real φ}) : + (mk_complex K φ : infinite_place K) = mk (φ : K →+* ℂ) := rfl + @[simp] -lemma mk_real.apply (φ : {φ : K →+* ℂ // complex_embedding.is_real φ}) (x : K) : - complex.abs (φ x) = mk_real K φ x := apply φ x +lemma mk_real.apply (φ : {φ : K →+* ℂ // complex_embedding.is_real φ}) (x : K) : + mk_real K φ x = complex.abs (φ x) := apply φ x @[simp] -lemma mk_complex.apply (φ : {φ : K →+* ℂ // ¬ complex_embedding.is_real φ}) (x : K) : - complex.abs (φ x) = mk_complex K φ x := apply φ x +lemma mk_complex.apply (φ : {φ : K →+* ℂ // ¬ complex_embedding.is_real φ}) (x : K) : + mk_complex K φ x = complex.abs (φ x) := apply φ x variable [number_field K] @@ -415,7 +446,7 @@ begin { ext, simp only [finset.mem_subtype, finset.mem_univ, not_is_real_iff_is_complex], }, { ext w, rw [@finset.prod_congr _ _ _ _ _ (λ φ, w x) _ (eq.refl _) - (λ φ hφ, (mk_complex.apply K φ x).trans + (λ φ hφ, (mk_complex.apply K φ x).symm.trans (congr_fun (congr_arg coe_fn (finset.mem_filter.1 hφ).2) x)), finset.prod_const, mk_complex.filter_card K w], refl, }}}, From 950605e4b9b4e2827681f637ba997307814a5ca9 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 25 Feb 2023 14:27:10 +0000 Subject: [PATCH 0529/1291] chore(measure theory/group/measure): Generalize typeclasses (#18496) A few generalisations found by Alex. Co-authored-by: Alex J. Best --- src/measure_theory/group/measure.lean | 34 +++++++++++++++++---------- 1 file changed, 21 insertions(+), 13 deletions(-) diff --git a/src/measure_theory/group/measure.lean b/src/measure_theory/group/measure.lean index b2e6534481b58..a982613517d5b 100644 --- a/src/measure_theory/group/measure.lean +++ b/src/measure_theory/group/measure.lean @@ -210,16 +210,18 @@ end has_measurable_mul end mul -section group - -variables [group G] +section div_inv_monoid +variables [div_inv_monoid G] @[to_additive] lemma map_div_right_eq_self (μ : measure G) [is_mul_right_invariant μ] (g : G) : map (/ g) μ = μ := by simp_rw [div_eq_mul_inv, map_mul_right_eq_self μ g⁻¹] -variables [has_measurable_mul G] +end div_inv_monoid + +section group +variables [group G] [has_measurable_mul G] @[to_additive] lemma measure_preserving_div_right (μ : measure G) [is_mul_right_invariant μ] @@ -340,8 +342,8 @@ instance (μ : measure G) [sigma_finite μ] : sigma_finite μ.inv := end has_involutive_inv -section mul_inv -variables [group G] [has_measurable_mul G] [has_measurable_inv G] {μ : measure G} +section division_monoid +variables [division_monoid G] [has_measurable_mul G] [has_measurable_inv G] {μ : measure G} @[to_additive] instance [is_mul_left_invariant μ] : is_mul_right_invariant μ.inv := @@ -387,24 +389,30 @@ lemma map_mul_right_inv_eq_self (μ : measure G) [is_inv_invariant μ] [is_mul_l (g : G) : map (λ t, (g * t)⁻¹) μ = μ := (measure_preserving_mul_right_inv μ g).map_eq +end division_monoid + +section group +variables [group G] [has_measurable_mul G] [has_measurable_inv G] {μ : measure G} + @[to_additive] lemma map_div_left_ae (μ : measure G) [is_mul_left_invariant μ] [is_inv_invariant μ] (x : G) : filter.map (λ t, x / t) μ.ae = μ.ae := ((measurable_equiv.div_left x).map_ae μ).trans $ congr_arg ae $ map_div_left_eq_self μ x -end mul_inv +end group end measure section topological_group -variables [topological_space G] [borel_space G] {μ : measure G} -variables [group G] [topological_group G] +variables [topological_space G] [borel_space G] {μ : measure G} [group G] @[to_additive] -instance measure.regular.inv [t2_space G] [regular μ] : regular μ.inv := +instance measure.regular.inv [has_continuous_inv G] [t2_space G] [regular μ] : regular μ.inv := regular.map (homeomorph.inv G) +variables [topological_group G] + @[to_additive] lemma regular_inv_iff [t2_space G] : μ.inv.regular ↔ μ.regular := begin @@ -545,9 +553,9 @@ end end topological_group -section comm_group +section comm_semigroup -variables [comm_group G] +variables [comm_semigroup G] /-- In an abelian group every left invariant measure is also right-invariant. We don't declare the converse as an instance, since that would loop type-class inference, and @@ -561,7 +569,7 @@ instance is_mul_left_invariant.is_mul_right_invariant {μ : measure G} [is_mul_l ⟨λ g, by simp_rw [mul_comm, map_mul_left_eq_self]⟩ -end comm_group +end comm_semigroup section haar From 3f5c9d30716c775bda043456728a1a3ee31412e7 Mon Sep 17 00:00:00 2001 From: Devon Tuma Date: Sat, 25 Feb 2023 15:41:16 +0000 Subject: [PATCH 0530/1291] feat(probability/probability_mass_function): construct a `pmf` from a probability measure (#18270) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Given a measurable space `α` with a `measurable_singleton_class` instance, this PR defines a function `measure.to_pmf` that converts a probability measure `μ` to a `pmf`, by defining the mass of a point `x` to be the measure of the singleton set `{x}` under `μ`. --- src/measure_theory/measure/measure_space.lean | 10 ++ .../probability_mass_function/basic.lean | 96 +++++++++++++++---- .../probability_mass_function/monad.lean | 12 ++- 3 files changed, 101 insertions(+), 17 deletions(-) diff --git a/src/measure_theory/measure/measure_space.lean b/src/measure_theory/measure/measure_space.lean index b4388e38b58b2..5a7873f4399c1 100644 --- a/src/measure_theory/measure/measure_space.lean +++ b/src/measure_theory/measure/measure_space.lean @@ -1852,6 +1852,16 @@ end sum (λ a, μ {a} • dirac a) = μ := by simpa using (map_eq_sum μ id measurable_id).symm +/-- Given that `α` is a countable, measurable space with all singleton sets measurable, +write the measure of a set `s` as the sum of the measure of `{x}` for all `x ∈ s`. -/ +lemma tsum_indicator_apply_singleton [countable α] [measurable_singleton_class α] + (μ : measure α) (s : set α) (hs : measurable_set s) : + ∑' (x : α), s.indicator (λ x, μ {x}) x = μ s := +calc ∑' (x : α), s.indicator (λ x, μ {x}) x = measure.sum (λ a, μ {a} • measure.dirac a) s : + by simp only [measure.sum_apply _ hs, measure.smul_apply, smul_eq_mul, measure.dirac_apply, + set.indicator_apply, mul_ite, pi.one_apply, mul_one, mul_zero] + ... = μ s : by rw μ.sum_smul_dirac + omit m0 end sum diff --git a/src/probability/probability_mass_function/basic.lean b/src/probability/probability_mass_function/basic.lean index e6222b2afc63f..12b1c45f657ce 100644 --- a/src/probability/probability_mass_function/basic.lean +++ b/src/probability/probability_mass_function/basic.lean @@ -20,6 +20,9 @@ by assigning each set the sum of the probabilities of each of its elements. Under this outer measure, every set is Carathéodory-measurable, so we can further extend this to a `measure` on `α`, see `pmf.to_measure`. `pmf.to_measure.is_probability_measure` shows this associated measure is a probability measure. +Conversely, given a probability measure `μ` on a measurable space `α` with all singleton sets +measurable, `μ.to_pmf` constructs a `pmf` on `α`, setting the probability mass of a point `x` +to be the measure of the singleton set `{x}`. ## Tags @@ -27,7 +30,7 @@ probability mass function, discrete probability measure -/ noncomputable theory variables {α β γ : Type*} -open_locale classical big_operators nnreal ennreal +open_locale classical big_operators nnreal ennreal measure_theory /-- A probability mass function, or discrete probability measures is a function `α → ℝ≥0∞` such that the values have (infinite) sum `1`. -/ @@ -103,6 +106,13 @@ variables (p : pmf α) (s t : set α) lemma to_outer_measure_apply : p.to_outer_measure s = ∑' x, s.indicator p x := tsum_congr (λ x, smul_dirac_apply (p x) x s) +@[simp] lemma to_outer_measure_caratheodory : p.to_outer_measure.caratheodory = ⊤ := +begin + refine (eq_top_iff.2 $ le_trans (le_Inf $ λ x hx, _) (le_sum_caratheodory _)), + exact let ⟨y, hy⟩ := hx in ((le_of_eq (dirac_caratheodory y).symm).trans + (le_smul_caratheodory _ _)).trans (le_of_eq hy), +end + @[simp] lemma to_outer_measure_apply_finset (s : finset α) : p.to_outer_measure s = ∑ x in s, p x := begin @@ -118,6 +128,13 @@ begin { exact ite_eq_left_iff.2 (λ ha', false.elim $ ha' rfl) } end +lemma to_outer_measure_injective : (to_outer_measure : pmf α → outer_measure α).injective := +λ p q h, pmf.ext (λ x, (p.to_outer_measure_apply_singleton x).symm.trans + ((congr_fun (congr_arg _ h) _).trans $ q.to_outer_measure_apply_singleton x)) + +@[simp] lemma to_outer_measure_inj {p q : pmf α} : + p.to_outer_measure = q.to_outer_measure ↔ p = q := to_outer_measure_injective.eq_iff + lemma to_outer_measure_apply_eq_zero_iff : p.to_outer_measure s = 0 ↔ disjoint p.support s := begin rw [to_outer_measure_apply, ennreal.tsum_eq_zero], @@ -138,8 +155,7 @@ begin exact λ a ha, (p.apply_eq_zero_iff a).2 $ set.not_mem_subset h ha } end -@[simp] -lemma to_outer_measure_apply_inter_support : +@[simp] lemma to_outer_measure_apply_inter_support : p.to_outer_measure (s ∩ p.support) = p.to_outer_measure s := by simp only [to_outer_measure_apply, pmf.support, set.indicator_inter_support] @@ -157,15 +173,6 @@ le_antisymm (p.to_outer_measure_mono (h.symm ▸ (set.inter_subset_left t p.supp lemma to_outer_measure_apply_fintype [fintype α] : p.to_outer_measure s = ∑ x, s.indicator p x := (p.to_outer_measure_apply s).trans (tsum_eq_sum (λ x h, absurd (finset.mem_univ x) h)) -@[simp] -lemma to_outer_measure_caratheodory (p : pmf α) : (to_outer_measure p).caratheodory = ⊤ := -begin - refine (eq_top_iff.2 $ le_trans (le_Inf $ λ x hx, _) (le_sum_caratheodory _)), - obtain ⟨y, hy⟩ := hx, - exact ((le_of_eq (dirac_caratheodory y).symm).trans - (le_smul_caratheodory _ _)).trans (le_of_eq hy), -end - end outer_measure section measure @@ -191,8 +198,7 @@ lemma to_measure_apply (hs : measurable_set s) : p.to_measure s = ∑' x, s.indi lemma to_measure_apply_singleton (a : α) (h : measurable_set ({a} : set α)) : p.to_measure {a} = p a := -by simp [to_measure_apply_eq_to_outer_measure_apply p {a} h, - to_outer_measure_apply_singleton] +by simp [to_measure_apply_eq_to_outer_measure_apply _ _ h, to_outer_measure_apply_singleton] lemma to_measure_apply_eq_zero_iff (hs : measurable_set s) : p.to_measure s = 0 ↔ disjoint p.support s := @@ -223,6 +229,14 @@ section measurable_singleton_class variables [measurable_singleton_class α] +lemma to_measure_injective : (to_measure : pmf α → measure α).injective := +λ p q h, pmf.ext (λ x, (p.to_measure_apply_singleton x $ measurable_set_singleton x).symm.trans + ((congr_fun (congr_arg _ h) _).trans $ q.to_measure_apply_singleton x $ + measurable_set_singleton x)) + +@[simp] lemma to_measure_inj {p q : pmf α} : p.to_measure = q.to_measure ↔ p = q := +to_measure_injective.eq_iff + @[simp] lemma to_measure_apply_finset (s : finset α) : p.to_measure s = ∑ x in s, p x := (p.to_measure_apply_eq_to_outer_measure_apply s s.measurable_set).trans @@ -239,11 +253,61 @@ lemma to_measure_apply_fintype [fintype α] : p.to_measure s = ∑ x, s.indicato end measurable_singleton_class +end measure + +end pmf + +namespace measure_theory + +open pmf + +namespace measure + +/-- Given that `α` is a countable, measurable space with all singleton sets measurable, +we can convert any probability measure into a `pmf`, where the mass of a point +is the measure of the singleton set under the original measure. -/ +def to_pmf [countable α] [measurable_space α] [measurable_singleton_class α] + (μ : measure α) [h : is_probability_measure μ] : pmf α := +⟨λ x, μ ({x} : set α), ennreal.summable.has_sum_iff.2 (trans (symm $ +(tsum_indicator_apply_singleton μ set.univ measurable_set.univ).symm.trans + (tsum_congr (λ x, congr_fun (set.indicator_univ _) x))) (h.measure_univ))⟩ + +variables [countable α] [measurable_space α] [measurable_singleton_class α] + (μ : measure α) [is_probability_measure μ] + +lemma to_pmf_apply (x : α) : μ.to_pmf x = μ {x} := rfl + +@[simp] lemma to_pmf_to_measure : μ.to_pmf.to_measure = μ := +measure.ext (λ s hs, by simpa only [μ.to_pmf.to_measure_apply s hs, + ← μ.tsum_indicator_apply_singleton s hs]) + +end measure + +end measure_theory + +namespace pmf + +open measure_theory + /-- The measure associated to a `pmf` by `to_measure` is a probability measure -/ -instance to_measure.is_probability_measure (p : pmf α) : is_probability_measure (p.to_measure) := +instance to_measure.is_probability_measure [measurable_space α] (p : pmf α) : + is_probability_measure (p.to_measure) := ⟨by simpa only [measurable_set.univ, to_measure_apply_eq_to_outer_measure_apply, set.indicator_univ, to_outer_measure_apply, ennreal.coe_eq_one] using tsum_coe p⟩ -end measure +variables [countable α] [measurable_space α] [measurable_singleton_class α] + (p : pmf α) (μ : measure α) [is_probability_measure μ] + +@[simp] lemma to_measure_to_pmf : p.to_measure.to_pmf = p := +pmf.ext (λ x, by rw [← p.to_measure_apply_singleton x (measurable_set_singleton x), + p.to_measure.to_pmf_apply]) + +lemma to_measure_eq_iff_eq_to_pmf (μ : measure α) [is_probability_measure μ] : + p.to_measure = μ ↔ p = μ.to_pmf := +by rw [← to_measure_inj, measure.to_pmf_to_measure] + +lemma to_pmf_eq_iff_to_measure_eq (μ : measure α) [is_probability_measure μ] : + μ.to_pmf = p ↔ μ = p.to_measure := +by rw [← to_measure_inj, measure.to_pmf_to_measure] end pmf diff --git a/src/probability/probability_mass_function/monad.lean b/src/probability/probability_mass_function/monad.lean index 76cf6758ed6f0..042c4d0da2c0e 100644 --- a/src/probability/probability_mass_function/monad.lean +++ b/src/probability/probability_mass_function/monad.lean @@ -21,6 +21,7 @@ so that the second argument only needs to be defined on the support of the first noncomputable theory variables {α β γ : Type*} open_locale classical big_operators nnreal ennreal +open measure_theory namespace pmf @@ -54,11 +55,20 @@ begin exact ite_eq_right_iff.2 (λ hb, ite_eq_right_iff.2 (λ h, (ha $ h ▸ hb).elim)) } end +variable [measurable_space α] + /-- The measure of a set under `pure a` is `1` for sets containing `a` and `0` otherwise -/ -@[simp] lemma to_measure_pure_apply [measurable_space α] (hs : measurable_set s) : +@[simp] lemma to_measure_pure_apply (hs : measurable_set s) : (pure a).to_measure s = if a ∈ s then 1 else 0 := (to_measure_apply_eq_to_outer_measure_apply (pure a) s hs).trans (to_outer_measure_pure_apply a s) +lemma to_measure_pure : (pure a).to_measure = measure.dirac a := +measure.ext (λ s hs, by simpa only [to_measure_pure_apply a s hs, measure.dirac_apply' a hs]) + +@[simp] lemma to_pmf_dirac [countable α] [h : measurable_singleton_class α] : + (measure.dirac a).to_pmf = pure a := +by rw [to_pmf_eq_iff_to_measure_eq, to_measure_pure] + end measure end pure From 3353f3371120058977ce1e20bf7fc8986c0fb042 Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Sat, 25 Feb 2023 18:06:51 +0000 Subject: [PATCH 0531/1291] feat(analysis/fourier): Poisson summation (#18392) This PR proves Poisson's summation formula for continuous functions on the real line satisfying the following (somewhat inexplicit) hypothesis: for each compact $K$, the sum $\sum_{n \in \mathbb{Z}} \sup_{x \in K} \|f(x + n)\|$ converges. (In a future PR it will be shown that this is automatically satisfied when f has exponential decay, giving a less general but rather easier-to-apply result.) --- src/analysis/fourier/fourier_transform.lean | 4 +- src/analysis/fourier/poisson_summation.lean | 114 ++++++++++++++++++ .../fourier/riemann_lebesgue_lemma.lean | 2 +- 3 files changed, 118 insertions(+), 2 deletions(-) create mode 100644 src/analysis/fourier/poisson_summation.lean diff --git a/src/analysis/fourier/fourier_transform.lean b/src/analysis/fourier/fourier_transform.lean index 2393022800122..75287e8e09d99 100644 --- a/src/analysis/fourier/fourier_transform.lean +++ b/src/analysis/fourier/fourier_transform.lean @@ -255,10 +255,12 @@ lemma fourier_integral_def (f : ℝ → E) (w : ℝ) : fourier_integral f w = ∫ (v : ℝ), fourier_char [-(v * w)] • f v := rfl +localized "notation (name := fourier_integral) `𝓕` := real.fourier_integral" in fourier_transform + lemma fourier_integral_eq_integral_exp_smul {E : Type*} [normed_add_comm_group E] [complete_space E] [normed_space ℂ E] (f : ℝ → E) (w : ℝ) : - fourier_integral f w = ∫ (v : ℝ), complex.exp (↑(-2 * π * v * w) * complex.I) • f v := + 𝓕 f w = ∫ (v : ℝ), complex.exp (↑(-2 * π * v * w) * complex.I) • f v := by simp_rw [fourier_integral_def, real.fourier_char_apply, mul_neg, neg_mul, mul_assoc] end real diff --git a/src/analysis/fourier/poisson_summation.lean b/src/analysis/fourier/poisson_summation.lean new file mode 100644 index 0000000000000..53017d75ed73a --- /dev/null +++ b/src/analysis/fourier/poisson_summation.lean @@ -0,0 +1,114 @@ +/- +Copyright (c) 2023 David Loeffler. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: David Loeffler +-/ +import analysis.fourier.add_circle +import analysis.fourier.fourier_transform + +/-! +# Poisson's summation formula + +We prove Poisson's summation formula `∑ (n : ℤ), f n = ∑ (n : ℤ), 𝓕 f n`, where `𝓕 f` is the +Fourier transform of `f`, under the following hypotheses: +* `f` is a continuous function `ℝ → ℂ`. +* The sum `∑ (n : ℤ), 𝓕 f n` is convergent. +* For all compacts `K ⊂ ℝ`, the sum `∑ (n : ℤ), sup { ‖f(x + n)‖ | x ∈ K }` is convergent. + +## TODO + +* Show that the conditions on `f` are automatically satisfied for Schwartz functions. +-/ + +noncomputable theory + +open function (hiding comp_apply) complex real set (hiding restrict_apply) + topological_space filter measure_theory + +open_locale real big_operators filter fourier_transform + +local attribute [instance] real.fact_zero_lt_one + +open continuous_map + +/-- The key lemma for Poisson summation: the `m`-th Fourier coefficient of the periodic function +`∑' n : ℤ, f (x + n)` is the value at `m` of the Fourier transform of `f`. -/ +lemma real.fourier_coeff_tsum_comp_add {f : C(ℝ, ℂ)} + (hf : ∀ (K : compacts ℝ), summable (λ n : ℤ, ‖(f.comp (continuous_map.add_right n)).restrict K‖)) + (m : ℤ) : + fourier_coeff (periodic.lift $ f.periodic_tsum_comp_add_zsmul 1) m = 𝓕 f m := +begin + -- NB: This proof can be shortened somewhat by telescoping together some of the steps in the calc + -- block, but I think it's more legible this way. We start with preliminaries about the integrand. + let e : C(ℝ, ℂ) := (fourier (-m)).comp ⟨(coe : ℝ → unit_add_circle), continuous_quotient_mk⟩, + have neK : ∀ (K : compacts ℝ) (g : C(ℝ, ℂ)), ‖(e * g).restrict K‖ = ‖g.restrict K‖, + { have : ∀ (x : ℝ), ‖e x‖ = 1, from λ x, abs_coe_circle _, + intros K g, + simp_rw [norm_eq_supr_norm, restrict_apply, mul_apply, norm_mul, this, one_mul] }, + have eadd : ∀ (n : ℤ), e.comp (continuous_map.add_right n) = e, + { intro n, ext1 x, + have : periodic e 1, from periodic.comp (λ x, add_circle.coe_add_period 1 x) _, + simpa only [mul_one] using this.int_mul n x }, + -- Now the main argument. First unwind some definitions. + calc fourier_coeff (periodic.lift $ f.periodic_tsum_comp_add_zsmul 1) m + = ∫ x in 0..1, e x * (∑' n : ℤ, f.comp (continuous_map.add_right n)) x : + by simp_rw [fourier_coeff_eq_interval_integral _ m 0, div_one, one_smul, zero_add, comp_apply, + coe_mk, periodic.lift_coe, zsmul_one, smul_eq_mul] + -- Transform sum in C(ℝ, ℂ) evaluated at x into pointwise sum of values. + ... = ∫ x in 0..1, (∑' n : ℤ, (e * f.comp (continuous_map.add_right n)) x) : + by simp_rw [coe_mul, pi.mul_apply, ←tsum_apply (summable_of_locally_summable_norm hf), + tsum_mul_left] + -- Swap sum and integral. + ... = ∑' n : ℤ, ∫ x in 0..1, (e * f.comp (continuous_map.add_right n)) x : + begin + refine (interval_integral.tsum_interval_integral_eq_of_summable_norm _).symm, + convert hf ⟨uIcc 0 1, is_compact_uIcc⟩, + exact funext (λ n, neK _ _) + end + ... = ∑' n : ℤ, ∫ x in 0..1, (e * f).comp (continuous_map.add_right n) x : + begin + simp only [continuous_map.comp_apply, mul_comp] at eadd ⊢, + simp_rw eadd, + end + -- Rearrange sum of interval integrals into an integral over `ℝ`. + ... = ∫ x, e x * f x : + begin + suffices : integrable ⇑(e * f), from this.has_sum_interval_integral_comp_add_int.tsum_eq, + apply integrable_of_summable_norm_Icc, + convert hf ⟨Icc 0 1, is_compact_Icc⟩, + simp_rw [continuous_map.comp_apply, mul_comp] at eadd ⊢, + simp_rw eadd, + exact funext (λ n, neK ⟨Icc 0 1, is_compact_Icc⟩ _), + end + -- Minor tidying to finish + ... = 𝓕 f m : + begin + rw fourier_integral_eq_integral_exp_smul, + congr' 1 with x : 1, + rw [smul_eq_mul, comp_apply, coe_mk, fourier_coe_apply], + congr' 2, + push_cast, + ring + end +end + +/-- **Poisson's summation formula**. -/ +theorem real.tsum_eq_tsum_fourier_integral {f : C(ℝ, ℂ)} + (h_norm : ∀ (K : compacts ℝ), + summable (λ n : ℤ, ‖(f.comp $ continuous_map.add_right n).restrict K‖)) + (h_sum : summable (λ n : ℤ, 𝓕 f n)) : + ∑' (n : ℤ), f n = ∑' (n : ℤ), 𝓕 f n := +begin + let F : C(unit_add_circle, ℂ) := ⟨(f.periodic_tsum_comp_add_zsmul 1).lift, + continuous_coinduced_dom.mpr (map_continuous _)⟩, + have : summable (fourier_coeff F), + { convert h_sum, + exact funext (λ n, real.fourier_coeff_tsum_comp_add h_norm n) }, + convert (has_pointwise_sum_fourier_series_of_summable this 0).tsum_eq.symm using 1, + { have := (has_sum_apply (summable_of_locally_summable_norm h_norm).has_sum 0).tsum_eq, + simpa only [coe_mk, ←quotient_add_group.coe_zero, periodic.lift_coe, zsmul_one, comp_apply, + coe_add_right, zero_add] using this }, + { congr' 1 with n : 1, + rw [←real.fourier_coeff_tsum_comp_add h_norm n, fourier_eval_zero, smul_eq_mul, mul_one], + refl }, +end diff --git a/src/analysis/fourier/riemann_lebesgue_lemma.lean b/src/analysis/fourier/riemann_lebesgue_lemma.lean index 2b2d53dd53313..de839b44c1f68 100644 --- a/src/analysis/fourier/riemann_lebesgue_lemma.lean +++ b/src/analysis/fourier/riemann_lebesgue_lemma.lean @@ -170,7 +170,7 @@ open_locale fourier_transform tends to 0 at infinity. -/ lemma real.fourier_integral_zero_at_infty_of_continuous_compact_support (hc : continuous f) (hs : has_compact_support f) : - tendsto (real.fourier_integral f) (cocompact ℝ) (𝓝 0) := + tendsto (𝓕 f) (cocompact ℝ) (𝓝 0) := begin refine ((zero_at_infty_integral_mul_exp_of_continuous_compact_support hc hs).comp (tendsto_cocompact_mul_left₀ From c23aca359d4be6975b2a577d5cdb9d77b82c7407 Mon Sep 17 00:00:00 2001 From: Michael Stoll Date: Sat, 25 Feb 2023 19:13:42 +0000 Subject: [PATCH 0532/1291] feat(number_theory/pell_general): add general existence result for Pell's Equation (#18484) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This begins the development of the theory of Pell's Equation for general parameter `d` (a positive nonsquare integer). Note that the existing file `number_theory.pell` restricts to `d = a ^ 2 - 1` (which is sufficient for the application to prove Matiyasevich's theorem). This PR provides a proof of the existence of a nontrivial solution: ```lean theorem ex_nontriv_sol {d : ℤ} (h₀ : 0 < d) (hd : ¬ is_square d) : ∃ x y : ℤ, x ^ 2 - d * y ^ 2 = 1 ∧ y ≠ 0 ``` We follow the (standard) proof given in Ireland-Rosen, which uses Dirichlet's approximation theorem and multiple instances of the pigeonhole principle to obtain two distinct solutions of `x^2 - d*y^2 = m` (for some `m`) that are congruent mod `m`; a nontrivial solution of the original equation is then obtained by "dividing" one by the other. See [this Zulip thread](https://leanprover.zulipchat.com/#narrow/stream/217875-Is-there-code-for-X.3F/topic/Proving.20Pell's.20equation.20is.20solvable/near/322885832). --- docs/100.yaml | 8 +- docs/references.bib | 14 ++++ src/number_theory/pell_general.lean | 123 ++++++++++++++++++++++++++++ 3 files changed, 142 insertions(+), 3 deletions(-) create mode 100644 src/number_theory/pell_general.lean diff --git a/docs/100.yaml b/docs/100.yaml index 408d2f1487e0c..ecb9fd4cf2075 100644 --- a/docs/100.yaml +++ b/docs/100.yaml @@ -147,9 +147,11 @@ author : Yury G. Kudryashov 39: title : Solutions to Pell’s Equation - decl : pell.eq_pell - author : Mario Carneiro - note : "`d` is defined to be `a*a - 1` for an arbitrary `a > 1`." + decls : + - pell.eq_pell + - pell.exists_of_not_is_square + author : Mario Carneiro (first), Michael Stoll (second) + note : "In `pell.eq_pell`, `d` is defined to be `a*a - 1` for an arbitrary `a > 1`." 40: title : Minkowski’s Fundamental Theorem 41: diff --git a/docs/references.bib b/docs/references.bib index 9bd399e6dfc22..8840f08ce1fda 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -1258,6 +1258,20 @@ @Book{ iordanescu2003 zbl = {1073.17014} } +@Book{ IrelandRosen1990, + author = {Ireland, Kenneth and Rosen, Michael}, + title = {A classical introduction to modern number theory}, + series = {Graduate Texts in Mathematics}, + volume = {84}, + edition = {Second}, + publisher = {Springer-Verlag, New York}, + year = {1990}, + pages = {xiv+389}, + isbn = {0-387-97329-X}, + doi = {10.1007/978-1-4757-2103-4}, + url = {https://doi.org/10.1007/978-1-4757-2103-4} +} + @Article{ izhakian2016, title = {Supertropical quadratic forms I}, journal = {Journal of Pure and Applied Algebra}, diff --git a/src/number_theory/pell_general.lean b/src/number_theory/pell_general.lean new file mode 100644 index 0000000000000..d4c93e1d1f0a1 --- /dev/null +++ b/src/number_theory/pell_general.lean @@ -0,0 +1,123 @@ +/- +Copyright (c) 2023 Michael Stoll. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Michael Geißer, Michael Stoll +-/ + +import tactic.qify +import tactic.linear_combination +import data.zmod.basic +import number_theory.diophantine_approximation + +/-! +# Pell's Equation + +We prove the following + +**Theorem.** Let $d$ be a positive integer that is not a square. Then the equation +$x^2 - d y^2 = 1$ has a nontrivial (i.e., with $y \ne 0$) solution in integers. + +See `pell.exists_of_not_is_square`. + +This is the beginning of a development that aims at providing all of the essential theory +of Pell's Equation for general $d$ (as opposed to the contents of `number_theory.pell`, +which is specific to the case $d = a^2 - 1$ for some $a > 1$). + +## References + +* [K. Ireland, M. Rosen, *A classical introduction to modern number theory* + (Section 17.5)][IrelandRosen1990] + +## Tags + +Pell's equation +-/ + +namespace pell + +section existence + +open set real + +/-- If `d` is a positive integer that is not a square, then there is a nontrivial solution +to the Pell equation `x^2 - d*y^2 = 1`. -/ +theorem exists_of_not_is_square {d : ℤ} (h₀ : 0 < d) (hd : ¬ is_square d) : + ∃ x y : ℤ, x ^ 2 - d * y ^ 2 = 1 ∧ y ≠ 0 := +begin + let ξ : ℝ := sqrt d, + have hξ : irrational ξ, + { refine irrational_nrt_of_notint_nrt 2 d (sq_sqrt $ int.cast_nonneg.mpr h₀.le) _ two_pos, + rintro ⟨x, hx⟩, + refine hd ⟨x, @int.cast_injective ℝ _ _ d (x * x) _⟩, + rw [← sq_sqrt $ int.cast_nonneg.mpr h₀.le, int.cast_mul, ← hx, sq], }, + obtain ⟨M, hM₁⟩ := exists_int_gt (2 * |ξ| + 1), + have hM : {q : ℚ | |q.1 ^ 2 - d * q.2 ^ 2| < M}.infinite, + { refine infinite.mono (λ q h, _) (infinite_rat_abs_sub_lt_one_div_denom_sq_of_irrational hξ), + have h0 : 0 < (q.2 : ℝ) ^ 2 := pow_pos (nat.cast_pos.mpr q.pos) 2, + have h1 : (q.num : ℝ) / (q.denom : ℝ) = q := by exact_mod_cast q.num_div_denom, + rw [mem_set_of, abs_sub_comm, ← @int.cast_lt ℝ, ← div_lt_div_right (abs_pos_of_pos h0)], + push_cast, + rw [← abs_div, abs_sq, sub_div, mul_div_cancel _ h0.ne', + ← div_pow, h1, ← sq_sqrt (int.cast_pos.mpr h₀).le, sq_sub_sq, abs_mul, ← mul_one_div], + refine mul_lt_mul'' (((abs_add ξ q).trans _).trans_lt hM₁) h (abs_nonneg _) (abs_nonneg _), + rw [two_mul, add_assoc, add_le_add_iff_left, ← sub_le_iff_le_add'], + rw [mem_set_of, abs_sub_comm] at h, + refine (abs_sub_abs_le_abs_sub (q : ℝ) ξ).trans (h.le.trans _), + rw [div_le_one h0, one_le_sq_iff_one_le_abs, nat.abs_cast, nat.one_le_cast], + exact q.pos, }, + obtain ⟨m, hm⟩ : ∃ m : ℤ, {q : ℚ | q.1 ^ 2 - d * q.2 ^ 2 = m}.infinite, + { contrapose! hM, + simp only [not_infinite] at hM ⊢, + refine (congr_arg _ (ext (λ x, _))).mp (finite.bUnion (finite_Ioo (-M) M) (λ m _, hM m)), + simp only [abs_lt, mem_set_of_eq, mem_Ioo, mem_Union, exists_prop, exists_eq_right'], }, + have hm₀ : m ≠ 0, + { rintro rfl, + obtain ⟨q, hq⟩ := hm.nonempty, + rw [mem_set_of, sub_eq_zero, mul_comm] at hq, + obtain ⟨a, ha⟩ := (int.pow_dvd_pow_iff two_pos).mp ⟨d, hq⟩, + rw [ha, mul_pow, mul_right_inj' (pow_pos (int.coe_nat_pos.mpr q.pos) 2).ne'] at hq, + exact hd ⟨a, sq a ▸ hq.symm⟩, }, + haveI := ne_zero_iff.mpr (int.nat_abs_ne_zero.mpr hm₀), + let f : ℚ → (zmod m.nat_abs) × (zmod m.nat_abs) := λ q, (q.1, q.2), + obtain ⟨q₁, h₁ : q₁.1 ^ 2 - d * q₁.2 ^ 2 = m, q₂, h₂ : q₂.1 ^ 2 - d * q₂.2 ^ 2 = m, hne, hqf⟩ := + hm.exists_ne_map_eq_of_maps_to (maps_to_univ f _) finite_univ, + obtain ⟨hq1 : (q₁.1 : zmod m.nat_abs) = q₂.1, hq2 : (q₁.2 : zmod m.nat_abs) = q₂.2⟩ := + prod.ext_iff.mp hqf, + have hd₁ : m ∣ q₁.1 * q₂.1 - d * (q₁.2 * q₂.2), + { rw [← int.nat_abs_dvd, ← zmod.int_coe_zmod_eq_zero_iff_dvd], + push_cast, + rw [hq1, hq2, ← sq, ← sq], + norm_cast, + rw [zmod.int_coe_zmod_eq_zero_iff_dvd, int.nat_abs_dvd, nat.cast_pow, ← h₂], }, + have hd₂ : m ∣ q₁.1 * q₂.2 - q₂.1 * q₁.2, + { rw [← int.nat_abs_dvd, ← zmod.int_coe_eq_int_coe_iff_dvd_sub], + push_cast, + rw [hq1, hq2], }, + replace hm₀ : (m : ℚ) ≠ 0 := int.cast_ne_zero.mpr hm₀, + refine ⟨(q₁.1 * q₂.1 - d * (q₁.2 * q₂.2)) / m, (q₁.1 * q₂.2 - q₂.1 * q₁.2) / m, _, _⟩, + { qify [hd₁, hd₂], + field_simp [hm₀], + norm_cast, + conv_rhs {congr, rw sq, congr, rw ← h₁, skip, rw ← h₂}, + push_cast, + ring, }, + { qify [hd₂], + refine div_ne_zero_iff.mpr ⟨_, hm₀⟩, + exact_mod_cast mt sub_eq_zero.mp (mt rat.eq_iff_mul_eq_mul.mpr hne), }, +end + +/-- If `d` is a positive integer, then there is a nontrivial solution +to the Pell equation `x^2 - d*y^2 = 1` if and only if `d` is not a square. -/ +theorem exists_iff_not_is_square {d : ℤ} (h₀ : 0 < d) : + (∃ x y : ℤ, x ^ 2 - d * y ^ 2 = 1 ∧ y ≠ 0) ↔ ¬ is_square d := +begin + refine ⟨_, exists_of_not_is_square h₀⟩, + rintros ⟨x, y, hxy, hy⟩ ⟨a, rfl⟩, + rw [← sq, ← mul_pow, sq_sub_sq, int.mul_eq_one_iff_eq_one_or_neg_one] at hxy, + replace hxy := hxy.elim (λ h, h.1.trans h.2.symm) (λ h, h.1.trans h.2.symm), + simpa [mul_self_pos.mp h₀, sub_eq_add_neg, eq_neg_self_iff] using hxy, +end + +end existence + +end pell From e064a7bf82ad94c3c17b5128bbd860d1ec34874e Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Sun, 26 Feb 2023 01:11:55 +0000 Subject: [PATCH 0533/1291] feat(data/polynomial/div): prove that evaluation induces an isomorphism of algebras (#18480) Also prove that `aeval` coerces to `eval_ring_hom`. --- src/data/polynomial/algebra_map.lean | 7 +++++-- src/data/polynomial/div.lean | 20 ++++++++++++++++++++ src/data/polynomial/eval.lean | 2 ++ src/ring_theory/ideal/operations.lean | 14 ++++++++++++++ src/ring_theory/ideal/quotient.lean | 2 +- 5 files changed, 42 insertions(+), 3 deletions(-) diff --git a/src/data/polynomial/algebra_map.lean b/src/data/polynomial/algebra_map.lean index 5c196b62d056b..db400c57c4623 100644 --- a/src/data/polynomial/algebra_map.lean +++ b/src/data/polynomial/algebra_map.lean @@ -232,8 +232,11 @@ lemma aeval_algebra_map_apply_eq_algebra_map_eval (x : R) (p : R[X]) : aeval (algebra_map R A x) p = algebra_map R A (p.eval x) := aeval_alg_hom_apply (algebra.of_id R A) x p -@[simp] lemma coe_aeval_eq_eval (r : R) : - (aeval r : R[X] → R) = eval r := rfl +@[simp] lemma coe_aeval_eq_eval (r : R) : (aeval r : R[X] → R) = eval r := rfl + +@[simp] lemma coe_aeval_eq_eval_ring_hom (x : R) : + ((aeval x : R[X] →ₐ[R] R) : R[X] →+* R) = eval_ring_hom x := +rfl @[simp] lemma aeval_fn_apply {X : Type*} (g : R[X]) (f : X → R) (x : X) : ((aeval f) g) x = aeval (f x) g := diff --git a/src/data/polynomial/div.lean b/src/data/polynomial/div.lean index f221f4fe4bf94..14fbe38e2284a 100644 --- a/src/data/polynomial/div.lean +++ b/src/data/polynomial/div.lean @@ -3,6 +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, Johannes Hölzl, Scott Morrison, Jens Wagemaker -/ +import data.polynomial.algebra_map import data.polynomial.inductions import data.polynomial.monic import ring_theory.multiplicity @@ -430,6 +431,25 @@ end variable {R} +lemma ker_eval_ring_hom (x : R) : (eval_ring_hom x).ker = ideal.span {X - C x} := +by { ext y, simpa only [ideal.mem_span_singleton, dvd_iff_is_root] } + +/-- For a commutative ring $R$, evaluating a polynomial at an element $x \in R$ induces an +isomorphism of $R$-algebras $R[X] / \langle X - x \rangle \cong R$. -/ +noncomputable def quotient_span_X_sub_C_alg_equiv (x : R) : + (R[X] ⧸ ideal.span ({X - C x} : set R[X])) ≃ₐ[R] R := +(alg_equiv.restrict_scalars R $ ideal.quotient_equiv_alg_of_eq R + (by exact ker_eval_ring_hom x : ring_hom.ker (aeval x).to_ring_hom = _)).symm.trans $ + ideal.quotient_ker_alg_equiv_of_right_inverse $ λ _, eval_C + +@[simp] lemma quotient_span_X_sub_C_alg_equiv_mk (x : R) (p : R[X]) : + quotient_span_X_sub_C_alg_equiv x (ideal.quotient.mk _ p) = p.eval x := +rfl + +@[simp] lemma quotient_span_X_sub_C_alg_equiv_symm_apply (x : R) (y : R) : + (quotient_span_X_sub_C_alg_equiv x).symm y = algebra_map R _ y := +rfl + section multiplicity /-- An algorithm for deciding polynomial divisibility. The algorithm is "compute `p %ₘ q` and compare to `0`". diff --git a/src/data/polynomial/eval.lean b/src/data/polynomial/eval.lean index 28701b84de5a0..3a39fd1b3db4e 100644 --- a/src/data/polynomial/eval.lean +++ b/src/data/polynomial/eval.lean @@ -399,6 +399,8 @@ by rwa [is_root, eval, eval₂_eq_zero_of_dvd_of_eval₂_eq_zero _ _ hpq] lemma not_is_root_C (r a : R) (hr : r ≠ 0) : ¬ is_root (C r) a := by simpa using hr +lemma eval_surjective (x : R) : function.surjective $ eval x := λ y, ⟨C y, eval_C⟩ + end eval section comp diff --git a/src/ring_theory/ideal/operations.lean b/src/ring_theory/ideal/operations.lean index 313ce2f53cd4c..8456e3cd022bb 100644 --- a/src/ring_theory/ideal/operations.lean +++ b/src/ring_theory/ideal/operations.lean @@ -2170,6 +2170,20 @@ begin exact quotient.eq.mpr hab end +variable (R₁) + +/-- Quotienting by equal ideals gives equivalent algebras. -/ +def quotient_equiv_alg_of_eq {I J : ideal A} (h : I = J) : (A ⧸ I) ≃ₐ[R₁] A ⧸ J := +quotient_equiv_alg I J alg_equiv.refl $ h ▸ (map_id I).symm + +@[simp] lemma quotient_equiv_alg_of_eq_mk {I J : ideal A} (h : I = J) (x : A) : + quotient_equiv_alg_of_eq R₁ h (ideal.quotient.mk I x) = ideal.quotient.mk J x := +rfl + +@[simp] lemma quotient_equiv_alg_of_eq_symm {I J : ideal A} (h : I = J) : + (quotient_equiv_alg_of_eq R₁ h).symm = quotient_equiv_alg_of_eq R₁ h.symm := +by ext; refl + end quotient_algebra end comm_ring diff --git a/src/ring_theory/ideal/quotient.lean b/src/ring_theory/ideal/quotient.lean index ce7a83a3f0eae..2ebb7ee7072e5 100644 --- a/src/ring_theory/ideal/quotient.lean +++ b/src/ring_theory/ideal/quotient.lean @@ -231,7 +231,7 @@ end quotient /-- Quotienting by equal ideals gives equivalent rings. -See also `submodule.quot_equiv_of_eq`. +See also `submodule.quot_equiv_of_eq` and `ideal.quotient_equiv_alg_of_eq`. -/ def quot_equiv_of_eq {R : Type*} [comm_ring R] {I J : ideal R} (h : I = J) : (R ⧸ I) ≃+* R ⧸ J := From 1c857a1f6798cb054be942199463c2cf904cb937 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sun, 26 Feb 2023 10:03:03 +0000 Subject: [PATCH 0534/1291] chore(data/finset/lattice): Protect a few more lemmas (#18497) `finset.disjoint_sup_left`/`finset.disjoint_sup_right` conflict with `disjoint_sup_left`/`disjoint_sup_right` when `finset` is open, so we protect them. --- src/data/finset/lattice.lean | 4 ++-- src/order/sup_indep.lean | 2 +- 2 files changed, 3 insertions(+), 3 deletions(-) diff --git a/src/data/finset/lattice.lean b/src/data/finset/lattice.lean index 0a2be525db223..7b991d9957df4 100644 --- a/src/data/finset/lattice.lean +++ b/src/data/finset/lattice.lean @@ -423,10 +423,10 @@ lemma sup_inf_distrib_right (s : finset ι) (f : ι → α) (a : α) : s.sup f ⊓ a = s.sup (λ i, f i ⊓ a) := by { rw [_root_.inf_comm, s.sup_inf_distrib_left], simp_rw _root_.inf_comm } -lemma disjoint_sup_right : disjoint a (s.sup f) ↔ ∀ i ∈ s, disjoint a (f i) := +protected lemma disjoint_sup_right : disjoint a (s.sup f) ↔ ∀ i ∈ s, disjoint a (f i) := by simp only [disjoint_iff, sup_inf_distrib_left, sup_eq_bot_iff] -lemma disjoint_sup_left : disjoint (s.sup f) a ↔ ∀ i ∈ s, disjoint (f i) a := +protected lemma disjoint_sup_left : disjoint (s.sup f) a ↔ ∀ i ∈ s, disjoint (f i) a := by simp only [disjoint_iff, sup_inf_distrib_right, sup_eq_bot_iff] end order_bot diff --git a/src/order/sup_indep.lean b/src/order/sup_indep.lean index 487c9e623553d..a91dab0d8ad13 100644 --- a/src/order/sup_indep.lean +++ b/src/order/sup_indep.lean @@ -135,7 +135,7 @@ variables [distrib_lattice α] [order_bot α] {s : finset ι} {f : ι → α} lemma sup_indep_iff_pairwise_disjoint : s.sup_indep f ↔ (s : set ι).pairwise_disjoint f := ⟨sup_indep.pairwise_disjoint, λ hs t ht i hi hit, - disjoint_sup_right.2 $ λ j hj, hs hi (ht hj) (ne_of_mem_of_not_mem hj hit).symm⟩ + finset.disjoint_sup_right.2 $ λ j hj, hs hi (ht hj) (ne_of_mem_of_not_mem hj hit).symm⟩ alias sup_indep_iff_pairwise_disjoint ↔ sup_indep.pairwise_disjoint _root_.set.pairwise_disjoint.sup_indep From 497d1e06409995dd8ec95301fa8d8f3480187f4c Mon Sep 17 00:00:00 2001 From: negiizhao Date: Sun, 26 Feb 2023 13:55:42 +0000 Subject: [PATCH 0535/1291] docs(set_theory/lists): update docs (#18346) --- src/set_theory/lists.lean | 7 ++----- 1 file changed, 2 insertions(+), 5 deletions(-) diff --git a/src/set_theory/lists.lean b/src/set_theory/lists.lean index a9a067fbf96ee..c45a4fb7c6d2b 100644 --- a/src/set_theory/lists.lean +++ b/src/set_theory/lists.lean @@ -37,11 +37,8 @@ This calls for a two-steps definition of ZFA lists: * `lists' α tt`: Proper ZFA prelists. Defined inductively from the empty ZFA prelist (`lists'.nil`) and from appending a ZFA prelist to a proper ZFA prelist (`lists'.cons a l`). * `lists α`: ZFA lists. Sum of the atoms and proper ZFA prelists. - -## TODO - -The next step is to define ZFA sets as lists quotiented by `lists.equiv`. -(-/ +* `finsets`: ZFA sets. Defined as `lists` quotiented by `lists.equiv`, the extensional equivalence. +-/ variables {α : Type*} From 146a2eed7ad5887ade571e073d0805d2ac618043 Mon Sep 17 00:00:00 2001 From: Kevin Buzzard Date: Sun, 26 Feb 2023 13:55:43 +0000 Subject: [PATCH 0536/1291] fix (measure_theory/measure/measure_space_def): fix module docstring typos (#18500) --- src/measure_theory/measure/measure_space_def.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/measure_theory/measure/measure_space_def.lean b/src/measure_theory/measure/measure_space_def.lean index 569a74f515b09..c6379cbe713a8 100644 --- a/src/measure_theory/measure/measure_space_def.lean +++ b/src/measure_theory/measure/measure_space_def.lean @@ -16,10 +16,10 @@ Given a measurable space `α`, a measure on `α` is a function that sends measur extended nonnegative reals that satisfies the following conditions: 1. `μ ∅ = 0`; 2. `μ` is countably additive. This means that the measure of a countable union of pairwise disjoint - sets is equal to the measure of the individual sets. + sets is equal to the sum of the measures of the individual sets. Every measure can be canonically extended to an outer measure, so that it assigns values to -all subsets, not just the measurable subsets. On the other hand, a measure that is countably +all subsets, not just the measurable subsets. On the other hand, an outer measure that is countably additive on measurable sets can be restricted to measurable sets to obtain a measure. In this file a measure is defined to be an outer measure that is countably additive on measurable sets, with the additional assumption that the outer measure is the canonical From c6b53304a1c87e90e6c14e3b28ca8fdc39070217 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Sun, 26 Feb 2023 15:22:11 +0000 Subject: [PATCH 0537/1291] feat(set_theory/ordinal/fixed_point): add missing theorem (#18323) We add `apply_lt_nfp_bfamily` matching `apply_lt_nfp_family`, and rename the existing theorem with that name to `apply_lt_nfp_bfamily_iff`, matching `apply_lt_nfp_family_iff`. --- src/set_theory/ordinal/fixed_point.lean | 21 +++++++++++++-------- 1 file changed, 13 insertions(+), 8 deletions(-) diff --git a/src/set_theory/ordinal/fixed_point.lean b/src/set_theory/ordinal/fixed_point.lean index f0b504b61d1e8..3913b862d6c29 100644 --- a/src/set_theory/ordinal/fixed_point.lean +++ b/src/set_theory/ordinal/fixed_point.lean @@ -237,20 +237,25 @@ sup_le theorem nfp_bfamily_monotone (hf : ∀ i hi, monotone (f i hi)) : monotone (nfp_bfamily o f) := nfp_family_monotone (λ i, hf _ _) -theorem apply_lt_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 := +theorem apply_lt_nfp_bfamily (H : ∀ i hi, is_normal (f i hi)) {a b} (hb : b < nfp_bfamily o f a) + (i hi) : f i hi b < nfp_bfamily o f a := begin - unfold nfp_bfamily, - rw ←@apply_lt_nfp_family_iff _ (family_of_bfamily o f) (out_nonempty_iff_ne_zero.2 ho) - (λ i, H _ _), - refine ⟨λ h i, h _ (typein_lt_self i), λ h i hio, _⟩, rw ←family_of_bfamily_enum o f, - apply h + apply apply_lt_nfp_family _ hb, + exact λ _, H _ _ end +theorem apply_lt_nfp_bfamily_iff (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 := +⟨λ h, begin + haveI := out_nonempty_iff_ne_zero.2 ho, + refine (apply_lt_nfp_family_iff _).1 (λ _, h _ _), + exact λ _, H _ _, +end, apply_lt_nfp_bfamily H⟩ + theorem nfp_bfamily_le_apply (ho : o ≠ 0) (H : ∀ i hi, is_normal (f i hi)) {a b} : (∃ i hi, nfp_bfamily o f a ≤ f i hi b) ↔ nfp_bfamily o f a ≤ b := -by { rw ←not_iff_not, push_neg, convert apply_lt_nfp_bfamily ho H, simp only [not_le] } +by { rw ←not_iff_not, push_neg, convert apply_lt_nfp_bfamily_iff ho H, simp only [not_le] } theorem nfp_bfamily_le_fp (H : ∀ i hi, monotone (f i hi)) {a b} (ab : a ≤ b) (h : ∀ i hi, f i hi b ≤ b) : nfp_bfamily o f a ≤ b := From a66d07e27d5b5b8ac1147cacfe353478e5c14002 Mon Sep 17 00:00:00 2001 From: Michael Stoll Date: Sun, 26 Feb 2023 18:11:25 +0000 Subject: [PATCH 0538/1291] chore(number_theory/pell*): rename files, update doc (#18503) Following a suggestion by @ocfnash in a comment to [#18484](https://github.com/leanprover-community/mathlib/pull/18484), this * renames `number_theory.pell` to `number_theory.pell_matiyasevic` * renames `number_theory.pell_general` to `number_theory.pell` * updates documentation accordingly * moves the "TODO" note from `pell_matiyasevic` to `pell` (and updates it) See also [this Zulip thread](https://leanprover.zulipchat.com/#narrow/stream/217875-Is-there-code-for-X.3F/topic/Proving.20Pell's.20equation.20is.20solvable/near/322885832). --- src/number_theory/dioph.lean | 2 +- src/number_theory/pell.lean | 831 +++--------------------- src/number_theory/pell_general.lean | 123 ---- src/number_theory/pell_matiyasevic.lean | 770 ++++++++++++++++++++++ 4 files changed, 865 insertions(+), 861 deletions(-) delete mode 100644 src/number_theory/pell_general.lean create mode 100644 src/number_theory/pell_matiyasevic.lean diff --git a/src/number_theory/dioph.lean b/src/number_theory/dioph.lean index 151abc8723c37..5b5d2b23291d1 100644 --- a/src/number_theory/dioph.lean +++ b/src/number_theory/dioph.lean @@ -6,7 +6,7 @@ Authors: Mario Carneiro import data.fin.fin2 import data.pfun import data.vector3 -import number_theory.pell +import number_theory.pell_matiyasevic /-! # Diophantine functions and Matiyasevic's theorem diff --git a/src/number_theory/pell.lean b/src/number_theory/pell.lean index 5fed9482d43e1..6d4cccaab4c46 100644 --- a/src/number_theory/pell.lean +++ b/src/number_theory/pell.lean @@ -1,771 +1,128 @@ /- -Copyright (c) 2017 Mario Carneiro. All rights reserved. +Copyright (c) 2023 Michael Stoll. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Mario Carneiro +Authors: Michael Geißer, Michael Stoll -/ -import data.nat.modeq -import number_theory.zsqrtd.basic +import tactic.qify +import data.zmod.basic +import number_theory.diophantine_approximation /-! -# Pell's equation and Matiyasevic's theorem +# Pell's Equation -This file solves Pell's equation, i.e. integer solutions to `x ^ 2 - d * y ^ 2 = 1` in the special -case that `d = a ^ 2 - 1`. This is then applied to prove Matiyasevic's theorem that the power -function is Diophantine, which is the last key ingredient in the solution to Hilbert's tenth -problem. For the definition of Diophantine function, see `dioph.lean`. +We prove the following -## Main definition +**Theorem.** Let $d$ be a positive integer that is not a square. Then the equation +$x^2 - d y^2 = 1$ has a nontrivial (i.e., with $y \ne 0$) solution in integers. -* `pell` is a function assigning to a natural number `n` the `n`-th solution to Pell's equation - constructed recursively from the initial solution `(0, 1)`. +See `pell.exists_of_not_is_square`. -## Main statements - -* `eq_pell` shows that every solution to Pell's equation is recursively obtained using `pell` -* `matiyasevic` shows that a certain system of Diophantine equations has a solution if and only if - the first variable is the `x`-component in a solution to Pell's equation - the key step towards - Hilbert's tenth problem in Davis' version of Matiyasevic's theorem. -* `eq_pow_of_pell` shows that the power function is Diophantine. - -## Implementation notes - -The proof of Matiyasevic's theorem doesn't follow Matiyasevic's original account of using Fibonacci -numbers but instead Davis' variant of using solutions to Pell's equation. +This is the beginning of a development that aims at providing all of the essential theory +of Pell's Equation for general $d$ (as opposed to the contents of `number_theory.pell_matiyasevic`, +which is specific to the case $d = a^2 - 1$ for some $a > 1$). ## References -* [M. Carneiro, _A Lean formalization of Matiyasevič's theorem_][carneiro2018matiyasevic] -* [M. Davis, _Hilbert's tenth problem is unsolvable_][MR317916] +* [K. Ireland, M. Rosen, *A classical introduction to modern number theory* + (Section 17.5)][IrelandRosen1990] ## Tags -Pell's equation, Matiyasevic's theorem, Hilbert's tenth problem +Pell's equation ## TODO -* Provide solutions to Pell's equation for the case of arbitrary `d` (not just `d = a ^ 2 - 1` like - in the current version) and furthermore also for `x ^ 2 - d * y ^ 2 = -1`. +* Provide the structure theory of the solution set to Pell's equation + and furthermore also for `x ^ 2 - d * y ^ 2 = -1` and further generalizations. * Connect solutions to the continued fraction expansion of `√d`. -/ namespace pell -open nat - -section -parameters {a : ℕ} (a1 : 1 < a) - -include a1 -private def d := a*a - 1 - -@[simp] theorem d_pos : 0 < d := -tsub_pos_of_lt (mul_lt_mul a1 (le_of_lt a1) dec_trivial dec_trivial : 1*1 rw ← int.coe_nat_mul at h}; - exact int.coe_nat_inj h - -theorem yn_add (m n) : yn (m + n) = xn m * yn n + yn m * xn n := -by injection (pell_zd_add _ m n) with _ h; - repeat {rw ← int.coe_nat_add at h <|> rw ← int.coe_nat_mul at h}; - exact int.coe_nat_inj h -theorem pell_zd_sub {m n} (h : n ≤ m) : pell_zd (m - n) = pell_zd m * (pell_zd n).conj := -let t := pell_zd_add n (m - n) in -by rw [add_tsub_cancel_of_le h] at t; - rw [t, mul_comm (pell_zd _ n) _, mul_assoc, (is_pell_norm _).1 (is_pell_pell_zd _ _), mul_one] +section existence -theorem xz_sub {m n} (h : n ≤ m) : xz (m - n) = xz m * xz n - d * yz m * yz n := -by { rw [sub_eq_add_neg, ←mul_neg], exact congr_arg zsqrtd.re (pell_zd_sub a1 h) } +open set real -theorem yz_sub {m n} (h : n ≤ m) : yz (m - n) = xz n * yz m - xz m * yz n := -by { rw [sub_eq_add_neg, ←mul_neg, mul_comm, add_comm], - exact congr_arg zsqrtd.im (pell_zd_sub a1 h) } - -theorem xy_coprime (n) : (xn n).coprime (yn n) := -nat.coprime_of_dvd' $ λk kp kx ky, -let p := pell_eq n in by rw ← p; exact -nat.dvd_sub (le_of_lt $ nat.lt_of_sub_eq_succ p) - (kx.mul_left _) (ky.mul_left _) - -theorem strict_mono_y : strict_mono yn -| m 0 h := absurd h $ nat.not_lt_zero _ -| m (n+1) h := - have yn m ≤ yn n, from or.elim (lt_or_eq_of_le $ nat.le_of_succ_le_succ h) - (λhl, le_of_lt $ strict_mono_y hl) (λe, by rw e), - by simp; refine lt_of_le_of_lt _ (nat.lt_add_of_pos_left $ x_pos a1 n); - rw ← mul_one (yn a1 m); - exact mul_le_mul this (le_of_lt a1) (nat.zero_le _) (nat.zero_le _) - -theorem strict_mono_x : strict_mono xn -| m 0 h := absurd h $ nat.not_lt_zero _ -| m (n+1) h := - have xn m ≤ xn n, from or.elim (lt_or_eq_of_le $ nat.le_of_succ_le_succ h) - (λhl, le_of_lt $ strict_mono_x hl) (λe, by rw e), - by simp; refine lt_of_lt_of_le (lt_of_le_of_lt this _) (nat.le_add_right _ _); - have t := nat.mul_lt_mul_of_pos_left a1 (x_pos a1 n); rwa mul_one at t - -theorem yn_ge_n : Π n, n ≤ yn n -| 0 := nat.zero_le _ -| (n+1) := show n < yn (n+1), from lt_of_le_of_lt (yn_ge_n n) (strict_mono_y $ nat.lt_succ_self n) - -theorem y_mul_dvd (n) : ∀k, yn n ∣ yn (n * k) -| 0 := dvd_zero _ -| (k+1) := by rw [nat.mul_succ, yn_add]; exact - dvd_add (dvd_mul_left _ _) ((y_mul_dvd k).mul_right _) - -theorem y_dvd_iff (m n) : yn m ∣ yn n ↔ m ∣ n := -⟨λh, nat.dvd_of_mod_eq_zero $ (nat.eq_zero_or_pos _).resolve_right $ λhp, - have co : nat.coprime (yn m) (xn (m * (n / m))), from nat.coprime.symm $ - (xy_coprime _).coprime_dvd_right (y_mul_dvd m (n / m)), - have m0 : 0 < m, from m.eq_zero_or_pos.resolve_left $ - λe, by rw [e, nat.mod_zero] at hp; rw [e] at h; exact - ne_of_lt (strict_mono_y a1 hp) (eq_zero_of_zero_dvd h).symm, - by rw [← nat.mod_add_div n m, yn_add] at h; exact - not_le_of_gt (strict_mono_y _ $ nat.mod_lt n m0) - (nat.le_of_dvd (strict_mono_y _ hp) $ co.dvd_of_dvd_mul_right $ - (nat.dvd_add_iff_right $ (y_mul_dvd _ _ _).mul_left _).2 h), -λ⟨k, e⟩, by rw e; apply y_mul_dvd⟩ - -theorem xy_modeq_yn (n) : - ∀ k, xn (n * k) ≡ (xn n)^k [MOD (yn n)^2] - ∧ yn (n * k) ≡ k * (xn n)^(k-1) * yn n [MOD (yn n)^3] -| 0 := by constructor; simp -| (k+1) := - let ⟨hx, hy⟩ := xy_modeq_yn k in - have L : xn (n * k) * xn n + d * yn (n * k) * yn n ≡ xn n^k * xn n + 0 [MOD yn n^2], from - (hx.mul_right _ ).add $ modeq_zero_iff_dvd.2 $ - by rw pow_succ'; exact - mul_dvd_mul_right (dvd_mul_of_dvd_right (modeq_zero_iff_dvd.1 $ - (hy.of_dvd $ by simp [pow_succ']).trans $ modeq_zero_iff_dvd.2 $ - by simp [-mul_comm, -mul_assoc]) _) _, - have R : xn (n * k) * yn n + yn (n * k) * xn n ≡ - xn n^k * yn n + k * xn n^k * yn n [MOD yn n^3], from - modeq.add (by { rw pow_succ', exact hx.mul_right' _ }) $ - have k * xn n^(k - 1) * yn n * xn n = k * xn n^k * yn n, - by clear _let_match; cases k with k; simp [pow_succ', mul_comm, mul_left_comm], - by { rw ← this, exact hy.mul_right _ }, - by { rw [add_tsub_cancel_right, nat.mul_succ, xn_add, yn_add, pow_succ' (xn _ n), - nat.succ_mul, add_comm (k * xn _ n^k) (xn _ n^k), right_distrib], - exact ⟨L, R⟩ } - -theorem ysq_dvd_yy (n) : yn n * yn n ∣ yn (n * yn n) := -modeq_zero_iff_dvd.1 $ - ((xy_modeq_yn n (yn n)).right.of_dvd $ by simp [pow_succ]).trans - (modeq_zero_iff_dvd.2 $ by simp [mul_dvd_mul_left, mul_assoc]) - -theorem dvd_of_ysq_dvd {n t} (h : yn n * yn n ∣ yn t) : yn n ∣ t := -have nt : n ∣ t, from (y_dvd_iff n t).1 $ dvd_of_mul_left_dvd h, -n.eq_zero_or_pos.elim (λ n0, by rwa n0 at ⊢ nt) $ λ (n0l : 0 < n), -let ⟨k, ke⟩ := nt in -have yn n ∣ k * (xn n)^(k-1), from -nat.dvd_of_mul_dvd_mul_right (strict_mono_y n0l) $ modeq_zero_iff_dvd.1 $ - by have xm := (xy_modeq_yn a1 n k).right; rw ← ke at xm; exact - (xm.of_dvd $ by simp [pow_succ]).symm.trans h.modeq_zero_nat, -by rw ke; exact dvd_mul_of_dvd_right - (((xy_coprime _ _).pow_left _).symm.dvd_of_dvd_mul_right this) _ - -theorem pell_zd_succ_succ (n) : pell_zd (n + 2) + pell_zd n = (2 * a : ℕ) * pell_zd (n + 1) := -have (1:ℤ√d) + ⟨a, 1⟩ * ⟨a, 1⟩ = ⟨a, 1⟩ * (2 * a), -by { rw zsqrtd.coe_nat_val, change (⟨_,_⟩:ℤ√(d a1))=⟨_,_⟩, - rw dz_val, dsimp [az], rw zsqrtd.ext, dsimp, split; ring }, -by simpa [mul_add, mul_comm, mul_left_comm, add_comm] using congr_arg (* pell_zd a1 n) this - -theorem xy_succ_succ (n) : xn (n + 2) + xn n = (2 * a) * xn (n + 1) ∧ - yn (n + 2) + yn n = (2 * a) * yn (n + 1) := begin - have := pell_zd_succ_succ a1 n, unfold pell_zd at this, - erw [zsqrtd.smul_val (2 * a : ℕ)] at this, - injection this with h₁ h₂, - split; apply int.coe_nat_inj; [simpa using h₁, simpa using h₂] -end - -theorem xn_succ_succ (n) : xn (n + 2) + xn n = (2 * a) * xn (n + 1) := (xy_succ_succ n).1 -theorem yn_succ_succ (n) : yn (n + 2) + yn n = (2 * a) * yn (n + 1) := (xy_succ_succ n).2 - -theorem xz_succ_succ (n) : xz (n + 2) = (2 * a : ℕ) * xz (n + 1) - xz n := -eq_sub_of_add_eq $ by delta xz; rw [← int.coe_nat_add, ← int.coe_nat_mul, xn_succ_succ] - -theorem yz_succ_succ (n) : yz (n + 2) = (2 * a : ℕ) * yz (n + 1) - yz n := -eq_sub_of_add_eq $ by delta yz; rw [← int.coe_nat_add, ← int.coe_nat_mul, yn_succ_succ] - -theorem yn_modeq_a_sub_one : ∀ n, yn n ≡ n [MOD a-1] -| 0 := by simp -| 1 := by simp -| (n+2) := (yn_modeq_a_sub_one n).add_right_cancel $ - begin - rw [yn_succ_succ, (by ring : n + 2 + n = 2 * (n + 1))], - exact ((modeq_sub a1.le).mul_left 2).mul (yn_modeq_a_sub_one (n+1)), - end - -theorem yn_modeq_two : ∀ n, yn n ≡ n [MOD 2] -| 0 := by simp -| 1 := by simp -| (n+2) := (yn_modeq_two n).add_right_cancel $ - begin - rw [yn_succ_succ, mul_assoc, (by ring : n + 2 + n = 2 * (n + 1))], - exact (dvd_mul_right 2 _).modeq_zero_nat.trans (dvd_mul_right 2 _).zero_modeq_nat, - end - -section - -omit a1 -lemma x_sub_y_dvd_pow_lem (y2 y1 y0 yn1 yn0 xn1 xn0 ay a2 : ℤ) : - (a2 * yn1 - yn0) * ay + y2 - (a2 * xn1 - xn0) = - y2 - a2 * y1 + y0 + a2 * (yn1 * ay + y1 - xn1) - (yn0 * ay + y0 - xn0) := by ring - -end - -theorem x_sub_y_dvd_pow (y : ℕ) : - ∀ n, (2*a*y - y*y - 1 : ℤ) ∣ yz n * (a - y) + ↑(y^n) - xz n -| 0 := by simp [xz, yz, int.coe_nat_zero, int.coe_nat_one] -| 1 := by simp [xz, yz, int.coe_nat_zero, int.coe_nat_one] -| (n+2) := - have (2*a*y - y*y - 1 : ℤ) ∣ ↑(y^(n + 2)) - ↑(2 * a) * ↑(y^(n + 1)) + ↑(y^n), from - ⟨-↑(y^n), by { simp [pow_succ, mul_add, int.coe_nat_mul, - show ((2:ℕ):ℤ) = 2, from rfl, mul_comm, mul_left_comm], ring }⟩, - by { rw [xz_succ_succ, yz_succ_succ, x_sub_y_dvd_pow_lem ↑(y^(n+2)) ↑(y^(n+1)) ↑(y^n)], - exact - dvd_sub (dvd_add this $ (x_sub_y_dvd_pow (n+1)).mul_left _) (x_sub_y_dvd_pow n) } - -theorem xn_modeq_x2n_add_lem (n j) : xn n ∣ d * yn n * (yn n * xn j) + xn j := -have h1 : d * yn n * (yn n * xn j) + xn j = (d * yn n * yn n + 1) * xn j, - by simp [add_mul, mul_assoc], -have h2 : d * yn n * yn n + 1 = xn n * xn n, by apply int.coe_nat_inj; - repeat {rw int.coe_nat_add <|> rw int.coe_nat_mul}; exact - add_eq_of_eq_sub' (eq.symm $ pell_eqz _ _), -by rw h2 at h1; rw [h1, mul_assoc]; exact dvd_mul_right _ _ - -theorem xn_modeq_x2n_add (n j) : xn (2 * n + j) + xn j ≡ 0 [MOD xn n] := +/-- If `d` is a positive integer that is not a square, then there is a nontrivial solution +to the Pell equation `x^2 - d*y^2 = 1`. -/ +theorem exists_of_not_is_square {d : ℤ} (h₀ : 0 < d) (hd : ¬ is_square d) : + ∃ x y : ℤ, x ^ 2 - d * y ^ 2 = 1 ∧ y ≠ 0 := begin - rw [two_mul, add_assoc, xn_add, add_assoc, ←zero_add 0], - refine (dvd_mul_right (xn a1 n) (xn a1 (n + j))).modeq_zero_nat.add _, - rw [yn_add, left_distrib, add_assoc, ←zero_add 0], - exact ((dvd_mul_right _ _).mul_left _).modeq_zero_nat.add - (xn_modeq_x2n_add_lem _ _ _).modeq_zero_nat, + let ξ : ℝ := sqrt d, + have hξ : irrational ξ, + { refine irrational_nrt_of_notint_nrt 2 d (sq_sqrt $ int.cast_nonneg.mpr h₀.le) _ two_pos, + rintro ⟨x, hx⟩, + refine hd ⟨x, @int.cast_injective ℝ _ _ d (x * x) _⟩, + rw [← sq_sqrt $ int.cast_nonneg.mpr h₀.le, int.cast_mul, ← hx, sq], }, + obtain ⟨M, hM₁⟩ := exists_int_gt (2 * |ξ| + 1), + have hM : {q : ℚ | |q.1 ^ 2 - d * q.2 ^ 2| < M}.infinite, + { refine infinite.mono (λ q h, _) (infinite_rat_abs_sub_lt_one_div_denom_sq_of_irrational hξ), + have h0 : 0 < (q.2 : ℝ) ^ 2 := pow_pos (nat.cast_pos.mpr q.pos) 2, + have h1 : (q.num : ℝ) / (q.denom : ℝ) = q := by exact_mod_cast q.num_div_denom, + rw [mem_set_of, abs_sub_comm, ← @int.cast_lt ℝ, ← div_lt_div_right (abs_pos_of_pos h0)], + push_cast, + rw [← abs_div, abs_sq, sub_div, mul_div_cancel _ h0.ne', + ← div_pow, h1, ← sq_sqrt (int.cast_pos.mpr h₀).le, sq_sub_sq, abs_mul, ← mul_one_div], + refine mul_lt_mul'' (((abs_add ξ q).trans _).trans_lt hM₁) h (abs_nonneg _) (abs_nonneg _), + rw [two_mul, add_assoc, add_le_add_iff_left, ← sub_le_iff_le_add'], + rw [mem_set_of, abs_sub_comm] at h, + refine (abs_sub_abs_le_abs_sub (q : ℝ) ξ).trans (h.le.trans _), + rw [div_le_one h0, one_le_sq_iff_one_le_abs, nat.abs_cast, nat.one_le_cast], + exact q.pos, }, + obtain ⟨m, hm⟩ : ∃ m : ℤ, {q : ℚ | q.1 ^ 2 - d * q.2 ^ 2 = m}.infinite, + { contrapose! hM, + simp only [not_infinite] at hM ⊢, + refine (congr_arg _ (ext (λ x, _))).mp (finite.bUnion (finite_Ioo (-M) M) (λ m _, hM m)), + simp only [abs_lt, mem_set_of_eq, mem_Ioo, mem_Union, exists_prop, exists_eq_right'], }, + have hm₀ : m ≠ 0, + { rintro rfl, + obtain ⟨q, hq⟩ := hm.nonempty, + rw [mem_set_of, sub_eq_zero, mul_comm] at hq, + obtain ⟨a, ha⟩ := (int.pow_dvd_pow_iff two_pos).mp ⟨d, hq⟩, + rw [ha, mul_pow, mul_right_inj' (pow_pos (int.coe_nat_pos.mpr q.pos) 2).ne'] at hq, + exact hd ⟨a, sq a ▸ hq.symm⟩, }, + haveI := ne_zero_iff.mpr (int.nat_abs_ne_zero.mpr hm₀), + let f : ℚ → (zmod m.nat_abs) × (zmod m.nat_abs) := λ q, (q.1, q.2), + obtain ⟨q₁, h₁ : q₁.1 ^ 2 - d * q₁.2 ^ 2 = m, q₂, h₂ : q₂.1 ^ 2 - d * q₂.2 ^ 2 = m, hne, hqf⟩ := + hm.exists_ne_map_eq_of_maps_to (maps_to_univ f _) finite_univ, + obtain ⟨hq1 : (q₁.1 : zmod m.nat_abs) = q₂.1, hq2 : (q₁.2 : zmod m.nat_abs) = q₂.2⟩ := + prod.ext_iff.mp hqf, + have hd₁ : m ∣ q₁.1 * q₂.1 - d * (q₁.2 * q₂.2), + { rw [← int.nat_abs_dvd, ← zmod.int_coe_zmod_eq_zero_iff_dvd], + push_cast, + rw [hq1, hq2, ← sq, ← sq], + norm_cast, + rw [zmod.int_coe_zmod_eq_zero_iff_dvd, int.nat_abs_dvd, nat.cast_pow, ← h₂], }, + have hd₂ : m ∣ q₁.1 * q₂.2 - q₂.1 * q₁.2, + { rw [← int.nat_abs_dvd, ← zmod.int_coe_eq_int_coe_iff_dvd_sub], + push_cast, + rw [hq1, hq2], }, + replace hm₀ : (m : ℚ) ≠ 0 := int.cast_ne_zero.mpr hm₀, + refine ⟨(q₁.1 * q₂.1 - d * (q₁.2 * q₂.2)) / m, (q₁.1 * q₂.2 - q₂.1 * q₁.2) / m, _, _⟩, + { qify [hd₁, hd₂], + field_simp [hm₀], + norm_cast, + conv_rhs {congr, rw sq, congr, rw ← h₁, skip, rw ← h₂}, + push_cast, + ring, }, + { qify [hd₂], + refine div_ne_zero_iff.mpr ⟨_, hm₀⟩, + exact_mod_cast mt sub_eq_zero.mp (mt rat.eq_iff_mul_eq_mul.mpr hne), }, end -lemma xn_modeq_x2n_sub_lem {n j} (h : j ≤ n) : xn (2 * n - j) + xn j ≡ 0 [MOD xn n] := -have h1 : xz n ∣ ↑d * yz n * yz (n - j) + xz j, - by rw [yz_sub _ h, mul_sub_left_distrib, sub_add_eq_add_sub]; exact -dvd_sub - (by delta xz; delta yz; - repeat {rw ← int.coe_nat_add <|> rw ← int.coe_nat_mul}; rw mul_comm (xn a1 j) (yn a1 n); - exact int.coe_nat_dvd.2 (xn_modeq_x2n_add_lem _ _ _)) - ((dvd_mul_right _ _).mul_left _), +/-- If `d` is a positive integer, then there is a nontrivial solution +to the Pell equation `x^2 - d*y^2 = 1` if and only if `d` is not a square. -/ +theorem exists_iff_not_is_square {d : ℤ} (h₀ : 0 < d) : + (∃ x y : ℤ, x ^ 2 - d * y ^ 2 = 1 ∧ y ≠ 0) ↔ ¬ is_square d := begin - rw [two_mul, add_tsub_assoc_of_le h, xn_add, add_assoc, ←zero_add 0], - exact (dvd_mul_right _ _).modeq_zero_nat.add - (int.coe_nat_dvd.1 $ by simpa [xz, yz] using h1).modeq_zero_nat, + refine ⟨_, exists_of_not_is_square h₀⟩, + rintros ⟨x, y, hxy, hy⟩ ⟨a, rfl⟩, + rw [← sq, ← mul_pow, sq_sub_sq, int.mul_eq_one_iff_eq_one_or_neg_one] at hxy, + replace hxy := hxy.elim (λ h, h.1.trans h.2.symm) (λ h, h.1.trans h.2.symm), + simpa [mul_self_pos.mp h₀, sub_eq_add_neg, eq_neg_self_iff] using hxy, end -theorem xn_modeq_x2n_sub {n j} (h : j ≤ 2 * n) : xn (2 * n - j) + xn j ≡ 0 [MOD xn n] := -(le_total j n).elim xn_modeq_x2n_sub_lem - (λjn, have 2 * n - j + j ≤ n + j, by rw [tsub_add_cancel_of_le h, two_mul]; - exact nat.add_le_add_left jn _, - let t := xn_modeq_x2n_sub_lem (nat.le_of_add_le_add_right this) in - by rwa [tsub_tsub_cancel_of_le h, add_comm] at t) - -theorem xn_modeq_x4n_add (n j) : xn (4 * n + j) ≡ xn j [MOD xn n] := -modeq.add_right_cancel' (xn (2 * n + j)) $ -by refine @modeq.trans _ _ 0 _ _ (by rw add_comm; exact (xn_modeq_x2n_add _ _ _).symm); - rw [show 4*n = 2*n + 2*n, from right_distrib 2 2 n, add_assoc]; apply xn_modeq_x2n_add - -theorem xn_modeq_x4n_sub {n j} (h : j ≤ 2 * n) : xn (4 * n - j) ≡ xn j [MOD xn n] := -have h' : j ≤ 2*n, from le_trans h (by rw nat.succ_mul; apply nat.le_add_left), -modeq.add_right_cancel' (xn (2 * n - j)) $ -by refine @modeq.trans _ _ 0 _ _ (by rw add_comm; exact (xn_modeq_x2n_sub _ h).symm); - rw [show 4*n = 2*n + 2*n, from right_distrib 2 2 n, add_tsub_assoc_of_le h']; - apply xn_modeq_x2n_add - -theorem eq_of_xn_modeq_lem1 {i n} : Π {j}, i < j → j < n → xn i % xn n < xn j % xn n -| 0 ij _ := absurd ij (nat.not_lt_zero _) -| (j+1) ij jn := - suffices xn j % xn n < xn (j + 1) % xn n, from - (lt_or_eq_of_le (nat.le_of_succ_le_succ ij)).elim - (λh, lt_trans (eq_of_xn_modeq_lem1 h (le_of_lt jn)) this) - (λh, by rw h; exact this), - by rw [nat.mod_eq_of_lt (strict_mono_x _ (nat.lt_of_succ_lt jn)), - nat.mod_eq_of_lt (strict_mono_x _ jn)]; - exact strict_mono_x _ (nat.lt_succ_self _) - -theorem eq_of_xn_modeq_lem2 {n} (h : 2 * xn n = xn (n + 1)) : a = 2 ∧ n = 0 := -by rw [xn_succ, mul_comm] at h; exact -have n = 0, from n.eq_zero_or_pos.resolve_right $ λnp, - ne_of_lt (lt_of_le_of_lt (nat.mul_le_mul_left _ a1) - (nat.lt_add_of_pos_right $ mul_pos (d_pos a1) (strict_mono_y a1 np))) h, -by cases this; simp at h; exact ⟨h.symm, rfl⟩ - -theorem eq_of_xn_modeq_lem3 {i n} (npos : 0 < n) : - Π {j}, i < j → j ≤ 2 * n → j ≠ n → ¬(a = 2 ∧ n = 1 ∧ i = 0 ∧ j = 2) → xn i % xn n < xn j % xn n -| 0 ij _ _ _ := absurd ij (nat.not_lt_zero _) -| (j+1) ij j2n jnn ntriv := - have lem2 : ∀k > n, k ≤ 2*n → (↑(xn k % xn n) : ℤ) = xn n - xn (2 * n - k), from λk kn k2n, - let k2nl := lt_of_add_lt_add_right $ show 2*n-k+k < n+k, by - {rw tsub_add_cancel_of_le, rw two_mul; exact (add_lt_add_left kn n), exact k2n } in - have xle : xn (2 * n - k) ≤ xn n, from le_of_lt $ strict_mono_x k2nl, - suffices xn k % xn n = xn n - xn (2 * n - k), by rw [this, int.coe_nat_sub xle], - by - { rw ← nat.mod_eq_of_lt (nat.sub_lt (x_pos a1 n) (x_pos a1 (2 * n - k))), - apply modeq.add_right_cancel' (xn a1 (2 * n - k)), - rw [tsub_add_cancel_of_le xle], - have t := xn_modeq_x2n_sub_lem a1 k2nl.le, - rw tsub_tsub_cancel_of_le k2n at t, - exact t.trans dvd_rfl.zero_modeq_nat }, - (lt_trichotomy j n).elim - (λ (jn : j < n), eq_of_xn_modeq_lem1 ij (lt_of_le_of_ne jn jnn)) $ λ o, o.elim - (λ (jn : j = n), by - { cases jn, - apply int.lt_of_coe_nat_lt_coe_nat, - rw [lem2 (n+1) (nat.lt_succ_self _) j2n, - show 2 * n - (n + 1) = n - 1, by rw[two_mul, tsub_add_eq_tsub_tsub, add_tsub_cancel_right]], - refine lt_sub_left_of_add_lt (int.coe_nat_lt_coe_nat_of_lt _), - cases (lt_or_eq_of_le $ nat.le_of_succ_le_succ ij) with lin ein, - { rw nat.mod_eq_of_lt (strict_mono_x _ lin), - have ll : xn a1 (n-1) + xn a1 (n-1) ≤ xn a1 n, - { rw [← two_mul, mul_comm, show xn a1 n = xn a1 (n-1+1), - by rw [tsub_add_cancel_of_le (succ_le_of_lt npos)], xn_succ], - exact le_trans (nat.mul_le_mul_left _ a1) (nat.le_add_right _ _) }, - have npm : (n-1).succ = n := nat.succ_pred_eq_of_pos npos, - have il : i ≤ n - 1, { apply nat.le_of_succ_le_succ, rw npm, exact lin }, - cases lt_or_eq_of_le il with ill ile, - { exact lt_of_lt_of_le (nat.add_lt_add_left (strict_mono_x a1 ill) _) ll }, - { rw ile, - apply lt_of_le_of_ne ll, - rw ← two_mul, - exact λe, ntriv $ - let ⟨a2, s1⟩ := @eq_of_xn_modeq_lem2 _ a1 (n-1) - (by rwa [tsub_add_cancel_of_le (succ_le_of_lt npos)]) in - have n1 : n = 1, from le_antisymm (tsub_eq_zero_iff_le.mp s1) npos, - by rw [ile, a2, n1]; exact ⟨rfl, rfl, rfl, rfl⟩ } }, - { rw [ein, nat.mod_self, add_zero], - exact strict_mono_x _ (nat.pred_lt npos.ne') } }) - (λ (jn : j > n), - have lem1 : j ≠ n → xn j % xn n < xn (j + 1) % xn n → xn i % xn n < xn (j + 1) % xn n, - from λjn s, - (lt_or_eq_of_le (nat.le_of_succ_le_succ ij)).elim - (λh, lt_trans (eq_of_xn_modeq_lem3 h (le_of_lt j2n) jn $ λ⟨a1, n1, i0, j2⟩, - by rw [n1, j2] at j2n; exact absurd j2n dec_trivial) s) - (λh, by rw h; exact s), - lem1 (ne_of_gt jn) $ int.lt_of_coe_nat_lt_coe_nat $ by - { rw [lem2 j jn (le_of_lt j2n), lem2 (j+1) (nat.le_succ_of_le jn) j2n], - refine sub_lt_sub_left (int.coe_nat_lt_coe_nat_of_lt $ strict_mono_x _ _) _, - rw [nat.sub_succ], - exact nat.pred_lt (ne_of_gt $ tsub_pos_of_lt j2n) }) - -theorem eq_of_xn_modeq_le {i j n} (ij : i ≤ j) (j2n : j ≤ 2 * n) - (h : xn i ≡ xn j [MOD xn n]) (ntriv : ¬(a = 2 ∧ n = 1 ∧ i = 0 ∧ j = 2)) : i = j := -if npos : n = 0 then by simp [*] at * else -(lt_or_eq_of_le ij).resolve_left $ λij', -if jn : j = n then by -{ refine ne_of_gt _ h, - rw [jn, nat.mod_self], - have x0 : 0 < xn a1 0 % xn a1 n := - by rw [nat.mod_eq_of_lt (strict_mono_x a1 (nat.pos_of_ne_zero npos))]; exact dec_trivial, - cases i with i, exact x0, - rw jn at ij', - exact x0.trans (eq_of_xn_modeq_lem3 _ (nat.pos_of_ne_zero npos) (nat.succ_pos _) - (le_trans ij j2n) (ne_of_lt ij') $ - λ⟨a1, n1, _, i2⟩, by rw [n1, i2] at ij'; exact absurd ij' dec_trivial) } -else ne_of_lt (eq_of_xn_modeq_lem3 (nat.pos_of_ne_zero npos) ij' j2n jn ntriv) h - -theorem eq_of_xn_modeq {i j n} (i2n : i ≤ 2 * n) (j2n : j ≤ 2 * n) - (h : xn i ≡ xn j [MOD xn n]) (ntriv : a = 2 → n = 1 → (i = 0 → j ≠ 2) ∧ (i = 2 → j ≠ 0)) : - i = j := -(le_total i j).elim - (λij, eq_of_xn_modeq_le ij j2n h $ λ⟨a2, n1, i0, j2⟩, (ntriv a2 n1).left i0 j2) - (λij, (eq_of_xn_modeq_le ij i2n h.symm $ λ⟨a2, n1, j0, i2⟩, - (ntriv a2 n1).right i2 j0).symm) - -theorem eq_of_xn_modeq' {i j n} (ipos : 0 < i) (hin : i ≤ n) (j4n : j ≤ 4 * n) - (h : xn j ≡ xn i [MOD xn n]) : j = i ∨ j + i = 4 * n := -have i2n : i ≤ 2*n, by apply le_trans hin; rw two_mul; apply nat.le_add_left, -(le_or_gt j (2 * n)).imp - (λj2n : j ≤ 2 * n, eq_of_xn_modeq j2n i2n h $ - λa2 n1, ⟨λj0 i2, by rw [n1, i2] at hin; exact absurd hin dec_trivial, - λj2 i0, ne_of_gt ipos i0⟩) - (λj2n : 2 * n < j, suffices i = 4*n - j, by rw [this, add_tsub_cancel_of_le j4n], - have j42n : 4*n - j ≤ 2*n, from @nat.le_of_add_le_add_right j _ _ $ - by rw [tsub_add_cancel_of_le j4n, show 4*n = 2*n + 2*n, from right_distrib 2 2 n]; - exact nat.add_le_add_left (le_of_lt j2n) _, - eq_of_xn_modeq i2n j42n - (h.symm.trans $ let t := xn_modeq_x4n_sub j42n in by rwa [tsub_tsub_cancel_of_le j4n] at t) - (λa2 n1, ⟨λi0, absurd i0 (ne_of_gt ipos), λi2, by { rw [n1, i2] at hin, - exact absurd hin dec_trivial }⟩)) - -theorem modeq_of_xn_modeq {i j n} (ipos : 0 < i) (hin : i ≤ n) (h : xn j ≡ xn i [MOD xn n]) : - j ≡ i [MOD 4 * n] ∨ j + i ≡ 0 [MOD 4 * n] := -let j' := j % (4 * n) in -have n4 : 0 < 4 * n, from mul_pos dec_trivial (ipos.trans_le hin), -have jl : j' < 4 * n, from nat.mod_lt _ n4, -have jj : j ≡ j' [MOD 4 * n], by delta modeq; rw nat.mod_eq_of_lt jl, -have ∀j q, xn (j + 4 * n * q) ≡ xn j [MOD xn n], begin - intros j q, induction q with q IH, { simp }, - rw [nat.mul_succ, ← add_assoc, add_comm], - exact (xn_modeq_x4n_add _ _ _).trans IH -end, -or.imp - (λ(ji : j' = i), by rwa ← ji) - (λ(ji : j' + i = 4 * n), (jj.add_right _).trans $ - by { rw ji, exact dvd_rfl.modeq_zero_nat }) - (eq_of_xn_modeq' ipos hin jl.le $ - (h.symm.trans $ by { rw ← nat.mod_add_div j (4*n), exact this j' _ }).symm) -end - -theorem xy_modeq_of_modeq {a b c} (a1 : 1 < a) (b1 : 1 < b) (h : a ≡ b [MOD c]) : - ∀ n, xn a1 n ≡ xn b1 n [MOD c] ∧ yn a1 n ≡ yn b1 n [MOD c] -| 0 := by constructor; refl -| 1 := by simp; exact ⟨h, modeq.refl 1⟩ -| (n+2) := ⟨ - (xy_modeq_of_modeq n).left.add_right_cancel $ - by { rw [xn_succ_succ a1, xn_succ_succ b1], exact - (h.mul_left _ ).mul (xy_modeq_of_modeq (n+1)).left }, - (xy_modeq_of_modeq n).right.add_right_cancel $ - by { rw [yn_succ_succ a1, yn_succ_succ b1], exact - (h.mul_left _ ).mul (xy_modeq_of_modeq (n+1)).right }⟩ - -theorem matiyasevic {a k x y} : (∃ a1 : 1 < a, xn a1 k = x ∧ yn a1 k = y) ↔ -1 < a ∧ k ≤ y ∧ -(x = 1 ∧ y = 0 ∨ -∃ (u v s t b : ℕ), - x * x - (a * a - 1) * y * y = 1 ∧ - u * u - (a * a - 1) * v * v = 1 ∧ - s * s - (b * b - 1) * t * t = 1 ∧ - 1 < b ∧ b ≡ 1 [MOD 4 * y] ∧ b ≡ a [MOD u] ∧ - 0 < v ∧ y * y ∣ v ∧ - s ≡ x [MOD u] ∧ - t ≡ k [MOD 4 * y]) := -⟨λ⟨a1, hx, hy⟩, by rw [← hx, ← hy]; - refine ⟨a1, (nat.eq_zero_or_pos k).elim - (λk0, by rw k0; exact ⟨le_rfl, or.inl ⟨rfl, rfl⟩⟩) (λkpos, _)⟩; exact - let x := xn a1 k, y := yn a1 k, - m := 2 * (k * y), - u := xn a1 m, v := yn a1 m in - have ky : k ≤ y, from yn_ge_n a1 k, - have yv : y * y ∣ v, from (ysq_dvd_yy a1 k).trans $ (y_dvd_iff _ _ _).2 $ dvd_mul_left _ _, - have uco : nat.coprime u (4 * y), from - have 2 ∣ v, from modeq_zero_iff_dvd.1 $ (yn_modeq_two _ _).trans - (dvd_mul_right _ _).modeq_zero_nat, - have nat.coprime u 2, from - (xy_coprime a1 m).coprime_dvd_right this, - (this.mul_right this).mul_right $ - (xy_coprime _ _).coprime_dvd_right (dvd_of_mul_left_dvd yv), - let ⟨b, ba, bm1⟩ := chinese_remainder uco a 1 in - have m1 : 1 < m, from - have 0 < k * y, from mul_pos kpos (strict_mono_y a1 kpos), - nat.mul_le_mul_left 2 this, - have vp : 0 < v, from strict_mono_y a1 (lt_trans zero_lt_one m1), - have b1 : 1 < b, from - have xn a1 1 < u, from strict_mono_x a1 m1, - have a < u, by simp at this; exact this, - lt_of_lt_of_le a1 $ by delta modeq at ba; - rw nat.mod_eq_of_lt this at ba; rw ← ba; apply nat.mod_le, - let s := xn b1 k, t := yn b1 k in - have sx : s ≡ x [MOD u], from (xy_modeq_of_modeq b1 a1 ba k).left, - have tk : t ≡ k [MOD 4 * y], from - have 4 * y ∣ b - 1, from int.coe_nat_dvd.1 $ - by rw int.coe_nat_sub (le_of_lt b1); - exact bm1.symm.dvd, - (yn_modeq_a_sub_one _ _).of_dvd this, - ⟨ky, or.inr ⟨u, v, s, t, b, - pell_eq _ _, pell_eq _ _, pell_eq _ _, b1, bm1, ba, vp, yv, sx, tk⟩⟩, -λ⟨a1, ky, o⟩, ⟨a1, match o with -| or.inl ⟨x1, y0⟩ := by rw y0 at ky; rw [nat.eq_zero_of_le_zero ky, x1, y0]; exact ⟨rfl, rfl⟩ -| or.inr ⟨u, v, s, t, b, xy, uv, st, b1, rem⟩ := - match x, y, eq_pell a1 xy, u, v, eq_pell a1 uv, s, t, eq_pell b1 st, rem, ky with - | ._, ._, ⟨i, rfl, rfl⟩, ._, ._, ⟨n, rfl, rfl⟩, ._, ._, ⟨j, rfl, rfl⟩, - ⟨(bm1 : b ≡ 1 [MOD 4 * yn a1 i]), - (ba : b ≡ a [MOD xn a1 n]), - (vp : 0 < yn a1 n), - (yv : yn a1 i * yn a1 i ∣ yn a1 n), - (sx : xn b1 j ≡ xn a1 i [MOD xn a1 n]), - (tk : yn b1 j ≡ k [MOD 4 * yn a1 i])⟩, - (ky : k ≤ yn a1 i) := - (nat.eq_zero_or_pos i).elim - (λi0, by simp [i0] at ky; rw [i0, ky]; exact ⟨rfl, rfl⟩) $ λipos, - suffices i = k, by rw this; exact ⟨rfl, rfl⟩, - by clear _x o rem xy uv st _match _match _fun_match; exact - have iln : i ≤ n, from le_of_not_gt $ λhin, - not_lt_of_ge (nat.le_of_dvd vp (dvd_of_mul_left_dvd yv)) (strict_mono_y a1 hin), - have yd : 4 * yn a1 i ∣ 4 * n, from mul_dvd_mul_left _ $ dvd_of_ysq_dvd a1 yv, - have jk : j ≡ k [MOD 4 * yn a1 i], from - have 4 * yn a1 i ∣ b - 1, from int.coe_nat_dvd.1 $ - by rw int.coe_nat_sub (le_of_lt b1); exact bm1.symm.dvd, - ((yn_modeq_a_sub_one b1 _).of_dvd this).symm.trans tk, - have ki : k + i < 4 * yn a1 i, from - lt_of_le_of_lt (add_le_add ky (yn_ge_n a1 i)) $ - by rw ← two_mul; exact nat.mul_lt_mul_of_pos_right dec_trivial (strict_mono_y a1 ipos), - have ji : j ≡ i [MOD 4 * n], from - have xn a1 j ≡ xn a1 i [MOD xn a1 n], from (xy_modeq_of_modeq b1 a1 ba j).left.symm.trans sx, - (modeq_of_xn_modeq a1 ipos iln this).resolve_right $ λ (ji : j + i ≡ 0 [MOD 4 * n]), - not_le_of_gt ki $ nat.le_of_dvd (lt_of_lt_of_le ipos $ nat.le_add_left _ _) $ - modeq_zero_iff_dvd.1 $ (jk.symm.add_right i).trans $ - ji.of_dvd yd, - by have : i % (4 * yn a1 i) = k % (4 * yn a1 i) := - (ji.of_dvd yd).symm.trans jk; - rwa [nat.mod_eq_of_lt (lt_of_le_of_lt (nat.le_add_left _ _) ki), - nat.mod_eq_of_lt (lt_of_le_of_lt (nat.le_add_right _ _) ki)] at this - end -end⟩⟩ - -lemma eq_pow_of_pell_lem {a y k} (hy0 : y ≠ 0) (hk0 : k ≠ 0) (hyk : y^k < a) : - (↑(y^k) : ℤ) < 2*a*y - y*y - 1 := -have hya : y < a, from (nat.le_self_pow hk0 _).trans_lt hyk, -calc (↑(y ^ k) : ℤ) < a : nat.cast_lt.2 hyk -... ≤ a ^ 2 - (a - 1) ^ 2 - 1 : - begin - rw [sub_sq, mul_one, one_pow, sub_add, sub_sub_cancel, two_mul, sub_sub, ← add_sub, - le_add_iff_nonneg_right, ← bit0, sub_nonneg, ← nat.cast_two, nat.cast_le, nat.succ_le_iff], - exact (one_le_iff_ne_zero.2 hy0).trans_lt hya - end -... ≤ a ^ 2 - (a - y) ^ 2 - 1 : have _ := hya.le, - by { mono*; simpa only [sub_nonneg, nat.cast_le, nat.one_le_cast, nat.one_le_iff_ne_zero] } -... = 2*a*y - y*y - 1 : by ring - -theorem eq_pow_of_pell {m n k} : n^k = m ↔ - k = 0 ∧ m = 1 ∨ - 0 < k ∧ (n = 0 ∧ m = 0 ∨ - 0 < n ∧ ∃ (w a t z : ℕ) (a1 : 1 < a), - xn a1 k ≡ yn a1 k * (a - n) + m [MOD t] ∧ - 2 * a * n = t + (n * n + 1) ∧ - m < t ∧ n ≤ w ∧ k ≤ w ∧ - a * a - ((w + 1) * (w + 1) - 1) * (w * z) * (w * z) = 1) := -begin - split, - { rintro rfl, - refine k.eq_zero_or_pos.imp (λ k0, k0.symm ▸ ⟨rfl, rfl⟩) (λ hk, ⟨hk, _⟩), - refine n.eq_zero_or_pos.imp (λ n0, n0.symm ▸ ⟨rfl, zero_pow hk⟩) (λ hn, ⟨hn, _⟩), - set w := max n k, - have nw : n ≤ w, from le_max_left _ _, - have kw : k ≤ w, from le_max_right _ _, - have wpos : 0 < w, from hn.trans_le nw, - have w1 : 1 < w + 1, from nat.succ_lt_succ wpos, - set a := xn w1 w, - have a1 : 1 < a, from strict_mono_x w1 wpos, - have na : n ≤ a, from nw.trans (n_lt_xn w1 w).le, - set x := xn a1 k, set y := yn a1 k, - obtain ⟨z, ze⟩ : w ∣ yn w1 w, - from modeq_zero_iff_dvd.1 ((yn_modeq_a_sub_one w1 w).trans dvd_rfl.modeq_zero_nat), - have nt : (↑(n^k) : ℤ) < 2 * a * n - n * n - 1, - { refine eq_pow_of_pell_lem hn.ne' hk.ne' _, - calc n^k ≤ n^w : nat.pow_le_pow_of_le_right hn kw - ... < (w + 1)^w : nat.pow_lt_pow_of_lt_left (nat.lt_succ_of_le nw) wpos - ... ≤ a : xn_ge_a_pow w1 w }, - lift (2 * a * n - n * n - 1 : ℤ) to ℕ using ((nat.cast_nonneg _).trans nt.le) with t te, - have tm : x ≡ y * (a - n) + n^k [MOD t], - { apply modeq_of_dvd, - rw [int.coe_nat_add, int.coe_nat_mul, int.coe_nat_sub na, te], - exact x_sub_y_dvd_pow a1 n k }, - have ta : 2 * a * n = t + (n * n + 1), - { rw [← @nat.cast_inj ℤ, int.coe_nat_add, te, sub_sub], - repeat { rw nat.cast_add <|> rw nat.cast_mul }, - rw [nat.cast_one, sub_add_cancel, nat.cast_two] }, - have zp : a * a - ((w + 1) * (w + 1) - 1) * (w * z) * (w * z) = 1, - from ze ▸ pell_eq w1 w, - exact ⟨w, a, t, z, a1, tm, ta, nat.cast_lt.1 nt, nw, kw, zp⟩ }, - { rintro (⟨rfl, rfl⟩ | ⟨hk0, ⟨rfl, rfl⟩ | ⟨hn0, w, a, t, z, a1, tm, ta, mt, nw, kw, zp⟩⟩), - { exact pow_zero n }, { exact zero_pow hk0 }, - have hw0 : 0 < w, from hn0.trans_le nw, - have hw1 : 1 < w + 1, from nat.succ_lt_succ hw0, - rcases eq_pell hw1 zp with ⟨j, rfl, yj⟩, - have hj0 : 0 < j, - { apply nat.pos_of_ne_zero, - rintro rfl, - exact lt_irrefl 1 a1 }, - have wj : w ≤ j := nat.le_of_dvd hj0 (modeq_zero_iff_dvd.1 $ - (yn_modeq_a_sub_one hw1 j).symm.trans $ modeq_zero_iff_dvd.2 ⟨z, yj.symm⟩), - have hnka : n ^ k < xn hw1 j, - calc n^k ≤ n^j : nat.pow_le_pow_of_le_right hn0 (le_trans kw wj) - ... < (w + 1)^j : nat.pow_lt_pow_of_lt_left (nat.lt_succ_of_le nw) hj0 - ... ≤ xn hw1 j : xn_ge_a_pow hw1 j, - have nt : (↑(n^k) : ℤ) < 2 * xn hw1 j * n - n * n - 1, - from eq_pow_of_pell_lem hn0.ne' hk0.ne' hnka, - have na : n ≤ xn hw1 j, from (nat.le_self_pow hk0.ne' _).trans hnka.le, - have te : (t : ℤ) = 2 * xn hw1 j * n - n * n - 1, - { rw [sub_sub, eq_sub_iff_add_eq], - exact_mod_cast ta.symm }, - have : xn a1 k ≡ yn a1 k * (xn hw1 j - n) + n^k [MOD t], - { apply modeq_of_dvd, - rw [te, nat.cast_add, nat.cast_mul, int.coe_nat_sub na], - exact x_sub_y_dvd_pow a1 n k }, - have : n^k % t = m % t, from (this.symm.trans tm).add_left_cancel' _, - rw [← te] at nt, - rwa [nat.mod_eq_of_lt (nat.cast_lt.1 nt), nat.mod_eq_of_lt mt] at this } -end +end existence end pell diff --git a/src/number_theory/pell_general.lean b/src/number_theory/pell_general.lean deleted file mode 100644 index d4c93e1d1f0a1..0000000000000 --- a/src/number_theory/pell_general.lean +++ /dev/null @@ -1,123 +0,0 @@ -/- -Copyright (c) 2023 Michael Stoll. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Michael Geißer, Michael Stoll --/ - -import tactic.qify -import tactic.linear_combination -import data.zmod.basic -import number_theory.diophantine_approximation - -/-! -# Pell's Equation - -We prove the following - -**Theorem.** Let $d$ be a positive integer that is not a square. Then the equation -$x^2 - d y^2 = 1$ has a nontrivial (i.e., with $y \ne 0$) solution in integers. - -See `pell.exists_of_not_is_square`. - -This is the beginning of a development that aims at providing all of the essential theory -of Pell's Equation for general $d$ (as opposed to the contents of `number_theory.pell`, -which is specific to the case $d = a^2 - 1$ for some $a > 1$). - -## References - -* [K. Ireland, M. Rosen, *A classical introduction to modern number theory* - (Section 17.5)][IrelandRosen1990] - -## Tags - -Pell's equation --/ - -namespace pell - -section existence - -open set real - -/-- If `d` is a positive integer that is not a square, then there is a nontrivial solution -to the Pell equation `x^2 - d*y^2 = 1`. -/ -theorem exists_of_not_is_square {d : ℤ} (h₀ : 0 < d) (hd : ¬ is_square d) : - ∃ x y : ℤ, x ^ 2 - d * y ^ 2 = 1 ∧ y ≠ 0 := -begin - let ξ : ℝ := sqrt d, - have hξ : irrational ξ, - { refine irrational_nrt_of_notint_nrt 2 d (sq_sqrt $ int.cast_nonneg.mpr h₀.le) _ two_pos, - rintro ⟨x, hx⟩, - refine hd ⟨x, @int.cast_injective ℝ _ _ d (x * x) _⟩, - rw [← sq_sqrt $ int.cast_nonneg.mpr h₀.le, int.cast_mul, ← hx, sq], }, - obtain ⟨M, hM₁⟩ := exists_int_gt (2 * |ξ| + 1), - have hM : {q : ℚ | |q.1 ^ 2 - d * q.2 ^ 2| < M}.infinite, - { refine infinite.mono (λ q h, _) (infinite_rat_abs_sub_lt_one_div_denom_sq_of_irrational hξ), - have h0 : 0 < (q.2 : ℝ) ^ 2 := pow_pos (nat.cast_pos.mpr q.pos) 2, - have h1 : (q.num : ℝ) / (q.denom : ℝ) = q := by exact_mod_cast q.num_div_denom, - rw [mem_set_of, abs_sub_comm, ← @int.cast_lt ℝ, ← div_lt_div_right (abs_pos_of_pos h0)], - push_cast, - rw [← abs_div, abs_sq, sub_div, mul_div_cancel _ h0.ne', - ← div_pow, h1, ← sq_sqrt (int.cast_pos.mpr h₀).le, sq_sub_sq, abs_mul, ← mul_one_div], - refine mul_lt_mul'' (((abs_add ξ q).trans _).trans_lt hM₁) h (abs_nonneg _) (abs_nonneg _), - rw [two_mul, add_assoc, add_le_add_iff_left, ← sub_le_iff_le_add'], - rw [mem_set_of, abs_sub_comm] at h, - refine (abs_sub_abs_le_abs_sub (q : ℝ) ξ).trans (h.le.trans _), - rw [div_le_one h0, one_le_sq_iff_one_le_abs, nat.abs_cast, nat.one_le_cast], - exact q.pos, }, - obtain ⟨m, hm⟩ : ∃ m : ℤ, {q : ℚ | q.1 ^ 2 - d * q.2 ^ 2 = m}.infinite, - { contrapose! hM, - simp only [not_infinite] at hM ⊢, - refine (congr_arg _ (ext (λ x, _))).mp (finite.bUnion (finite_Ioo (-M) M) (λ m _, hM m)), - simp only [abs_lt, mem_set_of_eq, mem_Ioo, mem_Union, exists_prop, exists_eq_right'], }, - have hm₀ : m ≠ 0, - { rintro rfl, - obtain ⟨q, hq⟩ := hm.nonempty, - rw [mem_set_of, sub_eq_zero, mul_comm] at hq, - obtain ⟨a, ha⟩ := (int.pow_dvd_pow_iff two_pos).mp ⟨d, hq⟩, - rw [ha, mul_pow, mul_right_inj' (pow_pos (int.coe_nat_pos.mpr q.pos) 2).ne'] at hq, - exact hd ⟨a, sq a ▸ hq.symm⟩, }, - haveI := ne_zero_iff.mpr (int.nat_abs_ne_zero.mpr hm₀), - let f : ℚ → (zmod m.nat_abs) × (zmod m.nat_abs) := λ q, (q.1, q.2), - obtain ⟨q₁, h₁ : q₁.1 ^ 2 - d * q₁.2 ^ 2 = m, q₂, h₂ : q₂.1 ^ 2 - d * q₂.2 ^ 2 = m, hne, hqf⟩ := - hm.exists_ne_map_eq_of_maps_to (maps_to_univ f _) finite_univ, - obtain ⟨hq1 : (q₁.1 : zmod m.nat_abs) = q₂.1, hq2 : (q₁.2 : zmod m.nat_abs) = q₂.2⟩ := - prod.ext_iff.mp hqf, - have hd₁ : m ∣ q₁.1 * q₂.1 - d * (q₁.2 * q₂.2), - { rw [← int.nat_abs_dvd, ← zmod.int_coe_zmod_eq_zero_iff_dvd], - push_cast, - rw [hq1, hq2, ← sq, ← sq], - norm_cast, - rw [zmod.int_coe_zmod_eq_zero_iff_dvd, int.nat_abs_dvd, nat.cast_pow, ← h₂], }, - have hd₂ : m ∣ q₁.1 * q₂.2 - q₂.1 * q₁.2, - { rw [← int.nat_abs_dvd, ← zmod.int_coe_eq_int_coe_iff_dvd_sub], - push_cast, - rw [hq1, hq2], }, - replace hm₀ : (m : ℚ) ≠ 0 := int.cast_ne_zero.mpr hm₀, - refine ⟨(q₁.1 * q₂.1 - d * (q₁.2 * q₂.2)) / m, (q₁.1 * q₂.2 - q₂.1 * q₁.2) / m, _, _⟩, - { qify [hd₁, hd₂], - field_simp [hm₀], - norm_cast, - conv_rhs {congr, rw sq, congr, rw ← h₁, skip, rw ← h₂}, - push_cast, - ring, }, - { qify [hd₂], - refine div_ne_zero_iff.mpr ⟨_, hm₀⟩, - exact_mod_cast mt sub_eq_zero.mp (mt rat.eq_iff_mul_eq_mul.mpr hne), }, -end - -/-- If `d` is a positive integer, then there is a nontrivial solution -to the Pell equation `x^2 - d*y^2 = 1` if and only if `d` is not a square. -/ -theorem exists_iff_not_is_square {d : ℤ} (h₀ : 0 < d) : - (∃ x y : ℤ, x ^ 2 - d * y ^ 2 = 1 ∧ y ≠ 0) ↔ ¬ is_square d := -begin - refine ⟨_, exists_of_not_is_square h₀⟩, - rintros ⟨x, y, hxy, hy⟩ ⟨a, rfl⟩, - rw [← sq, ← mul_pow, sq_sub_sq, int.mul_eq_one_iff_eq_one_or_neg_one] at hxy, - replace hxy := hxy.elim (λ h, h.1.trans h.2.symm) (λ h, h.1.trans h.2.symm), - simpa [mul_self_pos.mp h₀, sub_eq_add_neg, eq_neg_self_iff] using hxy, -end - -end existence - -end pell diff --git a/src/number_theory/pell_matiyasevic.lean b/src/number_theory/pell_matiyasevic.lean new file mode 100644 index 0000000000000..c36aeff75e744 --- /dev/null +++ b/src/number_theory/pell_matiyasevic.lean @@ -0,0 +1,770 @@ +/- +Copyright (c) 2017 Mario Carneiro. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Mario Carneiro +-/ + +import data.nat.modeq +import number_theory.zsqrtd.basic + +/-! +# Pell's equation and Matiyasevic's theorem + +This file solves Pell's equation, i.e. integer solutions to `x ^ 2 - d * y ^ 2 = 1` +*in the special case that `d = a ^ 2 - 1`*. +This is then applied to prove Matiyasevic's theorem that the power +function is Diophantine, which is the last key ingredient in the solution to Hilbert's tenth +problem. For the definition of Diophantine function, see `number_theory.dioph.lean`. + +For results on Pell's equation for arbitrary (positive, non-square) `d`, see +`number_theory.pell`. + +## Main definition + +* `pell` is a function assigning to a natural number `n` the `n`-th solution to Pell's equation + constructed recursively from the initial solution `(0, 1)`. + +## Main statements + +* `eq_pell` shows that every solution to Pell's equation is recursively obtained using `pell` +* `matiyasevic` shows that a certain system of Diophantine equations has a solution if and only if + the first variable is the `x`-component in a solution to Pell's equation - the key step towards + Hilbert's tenth problem in Davis' version of Matiyasevic's theorem. +* `eq_pow_of_pell` shows that the power function is Diophantine. + +## Implementation notes + +The proof of Matiyasevic's theorem doesn't follow Matiyasevic's original account of using Fibonacci +numbers but instead Davis' variant of using solutions to Pell's equation. + +## References + +* [M. Carneiro, _A Lean formalization of Matiyasevič's theorem_][carneiro2018matiyasevic] +* [M. Davis, _Hilbert's tenth problem is unsolvable_][MR317916] + +## Tags + +Pell's equation, Matiyasevic's theorem, Hilbert's tenth problem + +-/ + +namespace pell +open nat + +section +parameters {a : ℕ} (a1 : 1 < a) + +include a1 +private def d := a*a - 1 + +@[simp] theorem d_pos : 0 < d := +tsub_pos_of_lt (mul_lt_mul a1 (le_of_lt a1) dec_trivial dec_trivial : 1*1 rw ← int.coe_nat_mul at h}; + exact int.coe_nat_inj h + +theorem yn_add (m n) : yn (m + n) = xn m * yn n + yn m * xn n := +by injection (pell_zd_add _ m n) with _ h; + repeat {rw ← int.coe_nat_add at h <|> rw ← int.coe_nat_mul at h}; + exact int.coe_nat_inj h + +theorem pell_zd_sub {m n} (h : n ≤ m) : pell_zd (m - n) = pell_zd m * (pell_zd n).conj := +let t := pell_zd_add n (m - n) in +by rw [add_tsub_cancel_of_le h] at t; + rw [t, mul_comm (pell_zd _ n) _, mul_assoc, (is_pell_norm _).1 (is_pell_pell_zd _ _), mul_one] + +theorem xz_sub {m n} (h : n ≤ m) : xz (m - n) = xz m * xz n - d * yz m * yz n := +by { rw [sub_eq_add_neg, ←mul_neg], exact congr_arg zsqrtd.re (pell_zd_sub a1 h) } + +theorem yz_sub {m n} (h : n ≤ m) : yz (m - n) = xz n * yz m - xz m * yz n := +by { rw [sub_eq_add_neg, ←mul_neg, mul_comm, add_comm], + exact congr_arg zsqrtd.im (pell_zd_sub a1 h) } + +theorem xy_coprime (n) : (xn n).coprime (yn n) := +nat.coprime_of_dvd' $ λk kp kx ky, +let p := pell_eq n in by rw ← p; exact +nat.dvd_sub (le_of_lt $ nat.lt_of_sub_eq_succ p) + (kx.mul_left _) (ky.mul_left _) + +theorem strict_mono_y : strict_mono yn +| m 0 h := absurd h $ nat.not_lt_zero _ +| m (n+1) h := + have yn m ≤ yn n, from or.elim (lt_or_eq_of_le $ nat.le_of_succ_le_succ h) + (λhl, le_of_lt $ strict_mono_y hl) (λe, by rw e), + by simp; refine lt_of_le_of_lt _ (nat.lt_add_of_pos_left $ x_pos a1 n); + rw ← mul_one (yn a1 m); + exact mul_le_mul this (le_of_lt a1) (nat.zero_le _) (nat.zero_le _) + +theorem strict_mono_x : strict_mono xn +| m 0 h := absurd h $ nat.not_lt_zero _ +| m (n+1) h := + have xn m ≤ xn n, from or.elim (lt_or_eq_of_le $ nat.le_of_succ_le_succ h) + (λhl, le_of_lt $ strict_mono_x hl) (λe, by rw e), + by simp; refine lt_of_lt_of_le (lt_of_le_of_lt this _) (nat.le_add_right _ _); + have t := nat.mul_lt_mul_of_pos_left a1 (x_pos a1 n); rwa mul_one at t + +theorem yn_ge_n : Π n, n ≤ yn n +| 0 := nat.zero_le _ +| (n+1) := show n < yn (n+1), from lt_of_le_of_lt (yn_ge_n n) (strict_mono_y $ nat.lt_succ_self n) + +theorem y_mul_dvd (n) : ∀k, yn n ∣ yn (n * k) +| 0 := dvd_zero _ +| (k+1) := by rw [nat.mul_succ, yn_add]; exact + dvd_add (dvd_mul_left _ _) ((y_mul_dvd k).mul_right _) + +theorem y_dvd_iff (m n) : yn m ∣ yn n ↔ m ∣ n := +⟨λh, nat.dvd_of_mod_eq_zero $ (nat.eq_zero_or_pos _).resolve_right $ λhp, + have co : nat.coprime (yn m) (xn (m * (n / m))), from nat.coprime.symm $ + (xy_coprime _).coprime_dvd_right (y_mul_dvd m (n / m)), + have m0 : 0 < m, from m.eq_zero_or_pos.resolve_left $ + λe, by rw [e, nat.mod_zero] at hp; rw [e] at h; exact + ne_of_lt (strict_mono_y a1 hp) (eq_zero_of_zero_dvd h).symm, + by rw [← nat.mod_add_div n m, yn_add] at h; exact + not_le_of_gt (strict_mono_y _ $ nat.mod_lt n m0) + (nat.le_of_dvd (strict_mono_y _ hp) $ co.dvd_of_dvd_mul_right $ + (nat.dvd_add_iff_right $ (y_mul_dvd _ _ _).mul_left _).2 h), +λ⟨k, e⟩, by rw e; apply y_mul_dvd⟩ + +theorem xy_modeq_yn (n) : + ∀ k, xn (n * k) ≡ (xn n)^k [MOD (yn n)^2] + ∧ yn (n * k) ≡ k * (xn n)^(k-1) * yn n [MOD (yn n)^3] +| 0 := by constructor; simp +| (k+1) := + let ⟨hx, hy⟩ := xy_modeq_yn k in + have L : xn (n * k) * xn n + d * yn (n * k) * yn n ≡ xn n^k * xn n + 0 [MOD yn n^2], from + (hx.mul_right _ ).add $ modeq_zero_iff_dvd.2 $ + by rw pow_succ'; exact + mul_dvd_mul_right (dvd_mul_of_dvd_right (modeq_zero_iff_dvd.1 $ + (hy.of_dvd $ by simp [pow_succ']).trans $ modeq_zero_iff_dvd.2 $ + by simp [-mul_comm, -mul_assoc]) _) _, + have R : xn (n * k) * yn n + yn (n * k) * xn n ≡ + xn n^k * yn n + k * xn n^k * yn n [MOD yn n^3], from + modeq.add (by { rw pow_succ', exact hx.mul_right' _ }) $ + have k * xn n^(k - 1) * yn n * xn n = k * xn n^k * yn n, + by clear _let_match; cases k with k; simp [pow_succ', mul_comm, mul_left_comm], + by { rw ← this, exact hy.mul_right _ }, + by { rw [add_tsub_cancel_right, nat.mul_succ, xn_add, yn_add, pow_succ' (xn _ n), + nat.succ_mul, add_comm (k * xn _ n^k) (xn _ n^k), right_distrib], + exact ⟨L, R⟩ } + +theorem ysq_dvd_yy (n) : yn n * yn n ∣ yn (n * yn n) := +modeq_zero_iff_dvd.1 $ + ((xy_modeq_yn n (yn n)).right.of_dvd $ by simp [pow_succ]).trans + (modeq_zero_iff_dvd.2 $ by simp [mul_dvd_mul_left, mul_assoc]) + +theorem dvd_of_ysq_dvd {n t} (h : yn n * yn n ∣ yn t) : yn n ∣ t := +have nt : n ∣ t, from (y_dvd_iff n t).1 $ dvd_of_mul_left_dvd h, +n.eq_zero_or_pos.elim (λ n0, by rwa n0 at ⊢ nt) $ λ (n0l : 0 < n), +let ⟨k, ke⟩ := nt in +have yn n ∣ k * (xn n)^(k-1), from +nat.dvd_of_mul_dvd_mul_right (strict_mono_y n0l) $ modeq_zero_iff_dvd.1 $ + by have xm := (xy_modeq_yn a1 n k).right; rw ← ke at xm; exact + (xm.of_dvd $ by simp [pow_succ]).symm.trans h.modeq_zero_nat, +by rw ke; exact dvd_mul_of_dvd_right + (((xy_coprime _ _).pow_left _).symm.dvd_of_dvd_mul_right this) _ + +theorem pell_zd_succ_succ (n) : pell_zd (n + 2) + pell_zd n = (2 * a : ℕ) * pell_zd (n + 1) := +have (1:ℤ√d) + ⟨a, 1⟩ * ⟨a, 1⟩ = ⟨a, 1⟩ * (2 * a), +by { rw zsqrtd.coe_nat_val, change (⟨_,_⟩:ℤ√(d a1))=⟨_,_⟩, + rw dz_val, dsimp [az], rw zsqrtd.ext, dsimp, split; ring }, +by simpa [mul_add, mul_comm, mul_left_comm, add_comm] using congr_arg (* pell_zd a1 n) this + +theorem xy_succ_succ (n) : xn (n + 2) + xn n = (2 * a) * xn (n + 1) ∧ + yn (n + 2) + yn n = (2 * a) * yn (n + 1) := begin + have := pell_zd_succ_succ a1 n, unfold pell_zd at this, + erw [zsqrtd.smul_val (2 * a : ℕ)] at this, + injection this with h₁ h₂, + split; apply int.coe_nat_inj; [simpa using h₁, simpa using h₂] +end + +theorem xn_succ_succ (n) : xn (n + 2) + xn n = (2 * a) * xn (n + 1) := (xy_succ_succ n).1 +theorem yn_succ_succ (n) : yn (n + 2) + yn n = (2 * a) * yn (n + 1) := (xy_succ_succ n).2 + +theorem xz_succ_succ (n) : xz (n + 2) = (2 * a : ℕ) * xz (n + 1) - xz n := +eq_sub_of_add_eq $ by delta xz; rw [← int.coe_nat_add, ← int.coe_nat_mul, xn_succ_succ] + +theorem yz_succ_succ (n) : yz (n + 2) = (2 * a : ℕ) * yz (n + 1) - yz n := +eq_sub_of_add_eq $ by delta yz; rw [← int.coe_nat_add, ← int.coe_nat_mul, yn_succ_succ] + +theorem yn_modeq_a_sub_one : ∀ n, yn n ≡ n [MOD a-1] +| 0 := by simp +| 1 := by simp +| (n+2) := (yn_modeq_a_sub_one n).add_right_cancel $ + begin + rw [yn_succ_succ, (by ring : n + 2 + n = 2 * (n + 1))], + exact ((modeq_sub a1.le).mul_left 2).mul (yn_modeq_a_sub_one (n+1)), + end + +theorem yn_modeq_two : ∀ n, yn n ≡ n [MOD 2] +| 0 := by simp +| 1 := by simp +| (n+2) := (yn_modeq_two n).add_right_cancel $ + begin + rw [yn_succ_succ, mul_assoc, (by ring : n + 2 + n = 2 * (n + 1))], + exact (dvd_mul_right 2 _).modeq_zero_nat.trans (dvd_mul_right 2 _).zero_modeq_nat, + end + +section + +omit a1 +lemma x_sub_y_dvd_pow_lem (y2 y1 y0 yn1 yn0 xn1 xn0 ay a2 : ℤ) : + (a2 * yn1 - yn0) * ay + y2 - (a2 * xn1 - xn0) = + y2 - a2 * y1 + y0 + a2 * (yn1 * ay + y1 - xn1) - (yn0 * ay + y0 - xn0) := by ring + +end + +theorem x_sub_y_dvd_pow (y : ℕ) : + ∀ n, (2*a*y - y*y - 1 : ℤ) ∣ yz n * (a - y) + ↑(y^n) - xz n +| 0 := by simp [xz, yz, int.coe_nat_zero, int.coe_nat_one] +| 1 := by simp [xz, yz, int.coe_nat_zero, int.coe_nat_one] +| (n+2) := + have (2*a*y - y*y - 1 : ℤ) ∣ ↑(y^(n + 2)) - ↑(2 * a) * ↑(y^(n + 1)) + ↑(y^n), from + ⟨-↑(y^n), by { simp [pow_succ, mul_add, int.coe_nat_mul, + show ((2:ℕ):ℤ) = 2, from rfl, mul_comm, mul_left_comm], ring }⟩, + by { rw [xz_succ_succ, yz_succ_succ, x_sub_y_dvd_pow_lem ↑(y^(n+2)) ↑(y^(n+1)) ↑(y^n)], + exact + dvd_sub (dvd_add this $ (x_sub_y_dvd_pow (n+1)).mul_left _) (x_sub_y_dvd_pow n) } + +theorem xn_modeq_x2n_add_lem (n j) : xn n ∣ d * yn n * (yn n * xn j) + xn j := +have h1 : d * yn n * (yn n * xn j) + xn j = (d * yn n * yn n + 1) * xn j, + by simp [add_mul, mul_assoc], +have h2 : d * yn n * yn n + 1 = xn n * xn n, by apply int.coe_nat_inj; + repeat {rw int.coe_nat_add <|> rw int.coe_nat_mul}; exact + add_eq_of_eq_sub' (eq.symm $ pell_eqz _ _), +by rw h2 at h1; rw [h1, mul_assoc]; exact dvd_mul_right _ _ + +theorem xn_modeq_x2n_add (n j) : xn (2 * n + j) + xn j ≡ 0 [MOD xn n] := +begin + rw [two_mul, add_assoc, xn_add, add_assoc, ←zero_add 0], + refine (dvd_mul_right (xn a1 n) (xn a1 (n + j))).modeq_zero_nat.add _, + rw [yn_add, left_distrib, add_assoc, ←zero_add 0], + exact ((dvd_mul_right _ _).mul_left _).modeq_zero_nat.add + (xn_modeq_x2n_add_lem _ _ _).modeq_zero_nat, +end + +lemma xn_modeq_x2n_sub_lem {n j} (h : j ≤ n) : xn (2 * n - j) + xn j ≡ 0 [MOD xn n] := +have h1 : xz n ∣ ↑d * yz n * yz (n - j) + xz j, + by rw [yz_sub _ h, mul_sub_left_distrib, sub_add_eq_add_sub]; exact +dvd_sub + (by delta xz; delta yz; + repeat {rw ← int.coe_nat_add <|> rw ← int.coe_nat_mul}; rw mul_comm (xn a1 j) (yn a1 n); + exact int.coe_nat_dvd.2 (xn_modeq_x2n_add_lem _ _ _)) + ((dvd_mul_right _ _).mul_left _), +begin + rw [two_mul, add_tsub_assoc_of_le h, xn_add, add_assoc, ←zero_add 0], + exact (dvd_mul_right _ _).modeq_zero_nat.add + (int.coe_nat_dvd.1 $ by simpa [xz, yz] using h1).modeq_zero_nat, +end + +theorem xn_modeq_x2n_sub {n j} (h : j ≤ 2 * n) : xn (2 * n - j) + xn j ≡ 0 [MOD xn n] := +(le_total j n).elim xn_modeq_x2n_sub_lem + (λjn, have 2 * n - j + j ≤ n + j, by rw [tsub_add_cancel_of_le h, two_mul]; + exact nat.add_le_add_left jn _, + let t := xn_modeq_x2n_sub_lem (nat.le_of_add_le_add_right this) in + by rwa [tsub_tsub_cancel_of_le h, add_comm] at t) + +theorem xn_modeq_x4n_add (n j) : xn (4 * n + j) ≡ xn j [MOD xn n] := +modeq.add_right_cancel' (xn (2 * n + j)) $ +by refine @modeq.trans _ _ 0 _ _ (by rw add_comm; exact (xn_modeq_x2n_add _ _ _).symm); + rw [show 4*n = 2*n + 2*n, from right_distrib 2 2 n, add_assoc]; apply xn_modeq_x2n_add + +theorem xn_modeq_x4n_sub {n j} (h : j ≤ 2 * n) : xn (4 * n - j) ≡ xn j [MOD xn n] := +have h' : j ≤ 2*n, from le_trans h (by rw nat.succ_mul; apply nat.le_add_left), +modeq.add_right_cancel' (xn (2 * n - j)) $ +by refine @modeq.trans _ _ 0 _ _ (by rw add_comm; exact (xn_modeq_x2n_sub _ h).symm); + rw [show 4*n = 2*n + 2*n, from right_distrib 2 2 n, add_tsub_assoc_of_le h']; + apply xn_modeq_x2n_add + +theorem eq_of_xn_modeq_lem1 {i n} : Π {j}, i < j → j < n → xn i % xn n < xn j % xn n +| 0 ij _ := absurd ij (nat.not_lt_zero _) +| (j+1) ij jn := + suffices xn j % xn n < xn (j + 1) % xn n, from + (lt_or_eq_of_le (nat.le_of_succ_le_succ ij)).elim + (λh, lt_trans (eq_of_xn_modeq_lem1 h (le_of_lt jn)) this) + (λh, by rw h; exact this), + by rw [nat.mod_eq_of_lt (strict_mono_x _ (nat.lt_of_succ_lt jn)), + nat.mod_eq_of_lt (strict_mono_x _ jn)]; + exact strict_mono_x _ (nat.lt_succ_self _) + +theorem eq_of_xn_modeq_lem2 {n} (h : 2 * xn n = xn (n + 1)) : a = 2 ∧ n = 0 := +by rw [xn_succ, mul_comm] at h; exact +have n = 0, from n.eq_zero_or_pos.resolve_right $ λnp, + ne_of_lt (lt_of_le_of_lt (nat.mul_le_mul_left _ a1) + (nat.lt_add_of_pos_right $ mul_pos (d_pos a1) (strict_mono_y a1 np))) h, +by cases this; simp at h; exact ⟨h.symm, rfl⟩ + +theorem eq_of_xn_modeq_lem3 {i n} (npos : 0 < n) : + Π {j}, i < j → j ≤ 2 * n → j ≠ n → ¬(a = 2 ∧ n = 1 ∧ i = 0 ∧ j = 2) → xn i % xn n < xn j % xn n +| 0 ij _ _ _ := absurd ij (nat.not_lt_zero _) +| (j+1) ij j2n jnn ntriv := + have lem2 : ∀k > n, k ≤ 2*n → (↑(xn k % xn n) : ℤ) = xn n - xn (2 * n - k), from λk kn k2n, + let k2nl := lt_of_add_lt_add_right $ show 2*n-k+k < n+k, by + {rw tsub_add_cancel_of_le, rw two_mul; exact (add_lt_add_left kn n), exact k2n } in + have xle : xn (2 * n - k) ≤ xn n, from le_of_lt $ strict_mono_x k2nl, + suffices xn k % xn n = xn n - xn (2 * n - k), by rw [this, int.coe_nat_sub xle], + by + { rw ← nat.mod_eq_of_lt (nat.sub_lt (x_pos a1 n) (x_pos a1 (2 * n - k))), + apply modeq.add_right_cancel' (xn a1 (2 * n - k)), + rw [tsub_add_cancel_of_le xle], + have t := xn_modeq_x2n_sub_lem a1 k2nl.le, + rw tsub_tsub_cancel_of_le k2n at t, + exact t.trans dvd_rfl.zero_modeq_nat }, + (lt_trichotomy j n).elim + (λ (jn : j < n), eq_of_xn_modeq_lem1 ij (lt_of_le_of_ne jn jnn)) $ λ o, o.elim + (λ (jn : j = n), by + { cases jn, + apply int.lt_of_coe_nat_lt_coe_nat, + rw [lem2 (n+1) (nat.lt_succ_self _) j2n, + show 2 * n - (n + 1) = n - 1, by rw[two_mul, tsub_add_eq_tsub_tsub, add_tsub_cancel_right]], + refine lt_sub_left_of_add_lt (int.coe_nat_lt_coe_nat_of_lt _), + cases (lt_or_eq_of_le $ nat.le_of_succ_le_succ ij) with lin ein, + { rw nat.mod_eq_of_lt (strict_mono_x _ lin), + have ll : xn a1 (n-1) + xn a1 (n-1) ≤ xn a1 n, + { rw [← two_mul, mul_comm, show xn a1 n = xn a1 (n-1+1), + by rw [tsub_add_cancel_of_le (succ_le_of_lt npos)], xn_succ], + exact le_trans (nat.mul_le_mul_left _ a1) (nat.le_add_right _ _) }, + have npm : (n-1).succ = n := nat.succ_pred_eq_of_pos npos, + have il : i ≤ n - 1, { apply nat.le_of_succ_le_succ, rw npm, exact lin }, + cases lt_or_eq_of_le il with ill ile, + { exact lt_of_lt_of_le (nat.add_lt_add_left (strict_mono_x a1 ill) _) ll }, + { rw ile, + apply lt_of_le_of_ne ll, + rw ← two_mul, + exact λe, ntriv $ + let ⟨a2, s1⟩ := @eq_of_xn_modeq_lem2 _ a1 (n-1) + (by rwa [tsub_add_cancel_of_le (succ_le_of_lt npos)]) in + have n1 : n = 1, from le_antisymm (tsub_eq_zero_iff_le.mp s1) npos, + by rw [ile, a2, n1]; exact ⟨rfl, rfl, rfl, rfl⟩ } }, + { rw [ein, nat.mod_self, add_zero], + exact strict_mono_x _ (nat.pred_lt npos.ne') } }) + (λ (jn : j > n), + have lem1 : j ≠ n → xn j % xn n < xn (j + 1) % xn n → xn i % xn n < xn (j + 1) % xn n, + from λjn s, + (lt_or_eq_of_le (nat.le_of_succ_le_succ ij)).elim + (λh, lt_trans (eq_of_xn_modeq_lem3 h (le_of_lt j2n) jn $ λ⟨a1, n1, i0, j2⟩, + by rw [n1, j2] at j2n; exact absurd j2n dec_trivial) s) + (λh, by rw h; exact s), + lem1 (ne_of_gt jn) $ int.lt_of_coe_nat_lt_coe_nat $ by + { rw [lem2 j jn (le_of_lt j2n), lem2 (j+1) (nat.le_succ_of_le jn) j2n], + refine sub_lt_sub_left (int.coe_nat_lt_coe_nat_of_lt $ strict_mono_x _ _) _, + rw [nat.sub_succ], + exact nat.pred_lt (ne_of_gt $ tsub_pos_of_lt j2n) }) + +theorem eq_of_xn_modeq_le {i j n} (ij : i ≤ j) (j2n : j ≤ 2 * n) + (h : xn i ≡ xn j [MOD xn n]) (ntriv : ¬(a = 2 ∧ n = 1 ∧ i = 0 ∧ j = 2)) : i = j := +if npos : n = 0 then by simp [*] at * else +(lt_or_eq_of_le ij).resolve_left $ λij', +if jn : j = n then by +{ refine ne_of_gt _ h, + rw [jn, nat.mod_self], + have x0 : 0 < xn a1 0 % xn a1 n := + by rw [nat.mod_eq_of_lt (strict_mono_x a1 (nat.pos_of_ne_zero npos))]; exact dec_trivial, + cases i with i, exact x0, + rw jn at ij', + exact x0.trans (eq_of_xn_modeq_lem3 _ (nat.pos_of_ne_zero npos) (nat.succ_pos _) + (le_trans ij j2n) (ne_of_lt ij') $ + λ⟨a1, n1, _, i2⟩, by rw [n1, i2] at ij'; exact absurd ij' dec_trivial) } +else ne_of_lt (eq_of_xn_modeq_lem3 (nat.pos_of_ne_zero npos) ij' j2n jn ntriv) h + +theorem eq_of_xn_modeq {i j n} (i2n : i ≤ 2 * n) (j2n : j ≤ 2 * n) + (h : xn i ≡ xn j [MOD xn n]) (ntriv : a = 2 → n = 1 → (i = 0 → j ≠ 2) ∧ (i = 2 → j ≠ 0)) : + i = j := +(le_total i j).elim + (λij, eq_of_xn_modeq_le ij j2n h $ λ⟨a2, n1, i0, j2⟩, (ntriv a2 n1).left i0 j2) + (λij, (eq_of_xn_modeq_le ij i2n h.symm $ λ⟨a2, n1, j0, i2⟩, + (ntriv a2 n1).right i2 j0).symm) + +theorem eq_of_xn_modeq' {i j n} (ipos : 0 < i) (hin : i ≤ n) (j4n : j ≤ 4 * n) + (h : xn j ≡ xn i [MOD xn n]) : j = i ∨ j + i = 4 * n := +have i2n : i ≤ 2*n, by apply le_trans hin; rw two_mul; apply nat.le_add_left, +(le_or_gt j (2 * n)).imp + (λj2n : j ≤ 2 * n, eq_of_xn_modeq j2n i2n h $ + λa2 n1, ⟨λj0 i2, by rw [n1, i2] at hin; exact absurd hin dec_trivial, + λj2 i0, ne_of_gt ipos i0⟩) + (λj2n : 2 * n < j, suffices i = 4*n - j, by rw [this, add_tsub_cancel_of_le j4n], + have j42n : 4*n - j ≤ 2*n, from @nat.le_of_add_le_add_right j _ _ $ + by rw [tsub_add_cancel_of_le j4n, show 4*n = 2*n + 2*n, from right_distrib 2 2 n]; + exact nat.add_le_add_left (le_of_lt j2n) _, + eq_of_xn_modeq i2n j42n + (h.symm.trans $ let t := xn_modeq_x4n_sub j42n in by rwa [tsub_tsub_cancel_of_le j4n] at t) + (λa2 n1, ⟨λi0, absurd i0 (ne_of_gt ipos), λi2, by { rw [n1, i2] at hin, + exact absurd hin dec_trivial }⟩)) + +theorem modeq_of_xn_modeq {i j n} (ipos : 0 < i) (hin : i ≤ n) (h : xn j ≡ xn i [MOD xn n]) : + j ≡ i [MOD 4 * n] ∨ j + i ≡ 0 [MOD 4 * n] := +let j' := j % (4 * n) in +have n4 : 0 < 4 * n, from mul_pos dec_trivial (ipos.trans_le hin), +have jl : j' < 4 * n, from nat.mod_lt _ n4, +have jj : j ≡ j' [MOD 4 * n], by delta modeq; rw nat.mod_eq_of_lt jl, +have ∀j q, xn (j + 4 * n * q) ≡ xn j [MOD xn n], begin + intros j q, induction q with q IH, { simp }, + rw [nat.mul_succ, ← add_assoc, add_comm], + exact (xn_modeq_x4n_add _ _ _).trans IH +end, +or.imp + (λ(ji : j' = i), by rwa ← ji) + (λ(ji : j' + i = 4 * n), (jj.add_right _).trans $ + by { rw ji, exact dvd_rfl.modeq_zero_nat }) + (eq_of_xn_modeq' ipos hin jl.le $ + (h.symm.trans $ by { rw ← nat.mod_add_div j (4*n), exact this j' _ }).symm) +end + +theorem xy_modeq_of_modeq {a b c} (a1 : 1 < a) (b1 : 1 < b) (h : a ≡ b [MOD c]) : + ∀ n, xn a1 n ≡ xn b1 n [MOD c] ∧ yn a1 n ≡ yn b1 n [MOD c] +| 0 := by constructor; refl +| 1 := by simp; exact ⟨h, modeq.refl 1⟩ +| (n+2) := ⟨ + (xy_modeq_of_modeq n).left.add_right_cancel $ + by { rw [xn_succ_succ a1, xn_succ_succ b1], exact + (h.mul_left _ ).mul (xy_modeq_of_modeq (n+1)).left }, + (xy_modeq_of_modeq n).right.add_right_cancel $ + by { rw [yn_succ_succ a1, yn_succ_succ b1], exact + (h.mul_left _ ).mul (xy_modeq_of_modeq (n+1)).right }⟩ + +theorem matiyasevic {a k x y} : (∃ a1 : 1 < a, xn a1 k = x ∧ yn a1 k = y) ↔ +1 < a ∧ k ≤ y ∧ +(x = 1 ∧ y = 0 ∨ +∃ (u v s t b : ℕ), + x * x - (a * a - 1) * y * y = 1 ∧ + u * u - (a * a - 1) * v * v = 1 ∧ + s * s - (b * b - 1) * t * t = 1 ∧ + 1 < b ∧ b ≡ 1 [MOD 4 * y] ∧ b ≡ a [MOD u] ∧ + 0 < v ∧ y * y ∣ v ∧ + s ≡ x [MOD u] ∧ + t ≡ k [MOD 4 * y]) := +⟨λ⟨a1, hx, hy⟩, by rw [← hx, ← hy]; + refine ⟨a1, (nat.eq_zero_or_pos k).elim + (λk0, by rw k0; exact ⟨le_rfl, or.inl ⟨rfl, rfl⟩⟩) (λkpos, _)⟩; exact + let x := xn a1 k, y := yn a1 k, + m := 2 * (k * y), + u := xn a1 m, v := yn a1 m in + have ky : k ≤ y, from yn_ge_n a1 k, + have yv : y * y ∣ v, from (ysq_dvd_yy a1 k).trans $ (y_dvd_iff _ _ _).2 $ dvd_mul_left _ _, + have uco : nat.coprime u (4 * y), from + have 2 ∣ v, from modeq_zero_iff_dvd.1 $ (yn_modeq_two _ _).trans + (dvd_mul_right _ _).modeq_zero_nat, + have nat.coprime u 2, from + (xy_coprime a1 m).coprime_dvd_right this, + (this.mul_right this).mul_right $ + (xy_coprime _ _).coprime_dvd_right (dvd_of_mul_left_dvd yv), + let ⟨b, ba, bm1⟩ := chinese_remainder uco a 1 in + have m1 : 1 < m, from + have 0 < k * y, from mul_pos kpos (strict_mono_y a1 kpos), + nat.mul_le_mul_left 2 this, + have vp : 0 < v, from strict_mono_y a1 (lt_trans zero_lt_one m1), + have b1 : 1 < b, from + have xn a1 1 < u, from strict_mono_x a1 m1, + have a < u, by simp at this; exact this, + lt_of_lt_of_le a1 $ by delta modeq at ba; + rw nat.mod_eq_of_lt this at ba; rw ← ba; apply nat.mod_le, + let s := xn b1 k, t := yn b1 k in + have sx : s ≡ x [MOD u], from (xy_modeq_of_modeq b1 a1 ba k).left, + have tk : t ≡ k [MOD 4 * y], from + have 4 * y ∣ b - 1, from int.coe_nat_dvd.1 $ + by rw int.coe_nat_sub (le_of_lt b1); + exact bm1.symm.dvd, + (yn_modeq_a_sub_one _ _).of_dvd this, + ⟨ky, or.inr ⟨u, v, s, t, b, + pell_eq _ _, pell_eq _ _, pell_eq _ _, b1, bm1, ba, vp, yv, sx, tk⟩⟩, +λ⟨a1, ky, o⟩, ⟨a1, match o with +| or.inl ⟨x1, y0⟩ := by rw y0 at ky; rw [nat.eq_zero_of_le_zero ky, x1, y0]; exact ⟨rfl, rfl⟩ +| or.inr ⟨u, v, s, t, b, xy, uv, st, b1, rem⟩ := + match x, y, eq_pell a1 xy, u, v, eq_pell a1 uv, s, t, eq_pell b1 st, rem, ky with + | ._, ._, ⟨i, rfl, rfl⟩, ._, ._, ⟨n, rfl, rfl⟩, ._, ._, ⟨j, rfl, rfl⟩, + ⟨(bm1 : b ≡ 1 [MOD 4 * yn a1 i]), + (ba : b ≡ a [MOD xn a1 n]), + (vp : 0 < yn a1 n), + (yv : yn a1 i * yn a1 i ∣ yn a1 n), + (sx : xn b1 j ≡ xn a1 i [MOD xn a1 n]), + (tk : yn b1 j ≡ k [MOD 4 * yn a1 i])⟩, + (ky : k ≤ yn a1 i) := + (nat.eq_zero_or_pos i).elim + (λi0, by simp [i0] at ky; rw [i0, ky]; exact ⟨rfl, rfl⟩) $ λipos, + suffices i = k, by rw this; exact ⟨rfl, rfl⟩, + by clear _x o rem xy uv st _match _match _fun_match; exact + have iln : i ≤ n, from le_of_not_gt $ λhin, + not_lt_of_ge (nat.le_of_dvd vp (dvd_of_mul_left_dvd yv)) (strict_mono_y a1 hin), + have yd : 4 * yn a1 i ∣ 4 * n, from mul_dvd_mul_left _ $ dvd_of_ysq_dvd a1 yv, + have jk : j ≡ k [MOD 4 * yn a1 i], from + have 4 * yn a1 i ∣ b - 1, from int.coe_nat_dvd.1 $ + by rw int.coe_nat_sub (le_of_lt b1); exact bm1.symm.dvd, + ((yn_modeq_a_sub_one b1 _).of_dvd this).symm.trans tk, + have ki : k + i < 4 * yn a1 i, from + lt_of_le_of_lt (add_le_add ky (yn_ge_n a1 i)) $ + by rw ← two_mul; exact nat.mul_lt_mul_of_pos_right dec_trivial (strict_mono_y a1 ipos), + have ji : j ≡ i [MOD 4 * n], from + have xn a1 j ≡ xn a1 i [MOD xn a1 n], from (xy_modeq_of_modeq b1 a1 ba j).left.symm.trans sx, + (modeq_of_xn_modeq a1 ipos iln this).resolve_right $ λ (ji : j + i ≡ 0 [MOD 4 * n]), + not_le_of_gt ki $ nat.le_of_dvd (lt_of_lt_of_le ipos $ nat.le_add_left _ _) $ + modeq_zero_iff_dvd.1 $ (jk.symm.add_right i).trans $ + ji.of_dvd yd, + by have : i % (4 * yn a1 i) = k % (4 * yn a1 i) := + (ji.of_dvd yd).symm.trans jk; + rwa [nat.mod_eq_of_lt (lt_of_le_of_lt (nat.le_add_left _ _) ki), + nat.mod_eq_of_lt (lt_of_le_of_lt (nat.le_add_right _ _) ki)] at this + end +end⟩⟩ + +lemma eq_pow_of_pell_lem {a y k} (hy0 : y ≠ 0) (hk0 : k ≠ 0) (hyk : y^k < a) : + (↑(y^k) : ℤ) < 2*a*y - y*y - 1 := +have hya : y < a, from (nat.le_self_pow hk0 _).trans_lt hyk, +calc (↑(y ^ k) : ℤ) < a : nat.cast_lt.2 hyk +... ≤ a ^ 2 - (a - 1) ^ 2 - 1 : + begin + rw [sub_sq, mul_one, one_pow, sub_add, sub_sub_cancel, two_mul, sub_sub, ← add_sub, + le_add_iff_nonneg_right, ← bit0, sub_nonneg, ← nat.cast_two, nat.cast_le, nat.succ_le_iff], + exact (one_le_iff_ne_zero.2 hy0).trans_lt hya + end +... ≤ a ^ 2 - (a - y) ^ 2 - 1 : have _ := hya.le, + by { mono*; simpa only [sub_nonneg, nat.cast_le, nat.one_le_cast, nat.one_le_iff_ne_zero] } +... = 2*a*y - y*y - 1 : by ring + +theorem eq_pow_of_pell {m n k} : n^k = m ↔ + k = 0 ∧ m = 1 ∨ + 0 < k ∧ (n = 0 ∧ m = 0 ∨ + 0 < n ∧ ∃ (w a t z : ℕ) (a1 : 1 < a), + xn a1 k ≡ yn a1 k * (a - n) + m [MOD t] ∧ + 2 * a * n = t + (n * n + 1) ∧ + m < t ∧ n ≤ w ∧ k ≤ w ∧ + a * a - ((w + 1) * (w + 1) - 1) * (w * z) * (w * z) = 1) := +begin + split, + { rintro rfl, + refine k.eq_zero_or_pos.imp (λ k0, k0.symm ▸ ⟨rfl, rfl⟩) (λ hk, ⟨hk, _⟩), + refine n.eq_zero_or_pos.imp (λ n0, n0.symm ▸ ⟨rfl, zero_pow hk⟩) (λ hn, ⟨hn, _⟩), + set w := max n k, + have nw : n ≤ w, from le_max_left _ _, + have kw : k ≤ w, from le_max_right _ _, + have wpos : 0 < w, from hn.trans_le nw, + have w1 : 1 < w + 1, from nat.succ_lt_succ wpos, + set a := xn w1 w, + have a1 : 1 < a, from strict_mono_x w1 wpos, + have na : n ≤ a, from nw.trans (n_lt_xn w1 w).le, + set x := xn a1 k, set y := yn a1 k, + obtain ⟨z, ze⟩ : w ∣ yn w1 w, + from modeq_zero_iff_dvd.1 ((yn_modeq_a_sub_one w1 w).trans dvd_rfl.modeq_zero_nat), + have nt : (↑(n^k) : ℤ) < 2 * a * n - n * n - 1, + { refine eq_pow_of_pell_lem hn.ne' hk.ne' _, + calc n^k ≤ n^w : nat.pow_le_pow_of_le_right hn kw + ... < (w + 1)^w : nat.pow_lt_pow_of_lt_left (nat.lt_succ_of_le nw) wpos + ... ≤ a : xn_ge_a_pow w1 w }, + lift (2 * a * n - n * n - 1 : ℤ) to ℕ using ((nat.cast_nonneg _).trans nt.le) with t te, + have tm : x ≡ y * (a - n) + n^k [MOD t], + { apply modeq_of_dvd, + rw [int.coe_nat_add, int.coe_nat_mul, int.coe_nat_sub na, te], + exact x_sub_y_dvd_pow a1 n k }, + have ta : 2 * a * n = t + (n * n + 1), + { rw [← @nat.cast_inj ℤ, int.coe_nat_add, te, sub_sub], + repeat { rw nat.cast_add <|> rw nat.cast_mul }, + rw [nat.cast_one, sub_add_cancel, nat.cast_two] }, + have zp : a * a - ((w + 1) * (w + 1) - 1) * (w * z) * (w * z) = 1, + from ze ▸ pell_eq w1 w, + exact ⟨w, a, t, z, a1, tm, ta, nat.cast_lt.1 nt, nw, kw, zp⟩ }, + { rintro (⟨rfl, rfl⟩ | ⟨hk0, ⟨rfl, rfl⟩ | ⟨hn0, w, a, t, z, a1, tm, ta, mt, nw, kw, zp⟩⟩), + { exact pow_zero n }, { exact zero_pow hk0 }, + have hw0 : 0 < w, from hn0.trans_le nw, + have hw1 : 1 < w + 1, from nat.succ_lt_succ hw0, + rcases eq_pell hw1 zp with ⟨j, rfl, yj⟩, + have hj0 : 0 < j, + { apply nat.pos_of_ne_zero, + rintro rfl, + exact lt_irrefl 1 a1 }, + have wj : w ≤ j := nat.le_of_dvd hj0 (modeq_zero_iff_dvd.1 $ + (yn_modeq_a_sub_one hw1 j).symm.trans $ modeq_zero_iff_dvd.2 ⟨z, yj.symm⟩), + have hnka : n ^ k < xn hw1 j, + calc n^k ≤ n^j : nat.pow_le_pow_of_le_right hn0 (le_trans kw wj) + ... < (w + 1)^j : nat.pow_lt_pow_of_lt_left (nat.lt_succ_of_le nw) hj0 + ... ≤ xn hw1 j : xn_ge_a_pow hw1 j, + have nt : (↑(n^k) : ℤ) < 2 * xn hw1 j * n - n * n - 1, + from eq_pow_of_pell_lem hn0.ne' hk0.ne' hnka, + have na : n ≤ xn hw1 j, from (nat.le_self_pow hk0.ne' _).trans hnka.le, + have te : (t : ℤ) = 2 * xn hw1 j * n - n * n - 1, + { rw [sub_sub, eq_sub_iff_add_eq], + exact_mod_cast ta.symm }, + have : xn a1 k ≡ yn a1 k * (xn hw1 j - n) + n^k [MOD t], + { apply modeq_of_dvd, + rw [te, nat.cast_add, nat.cast_mul, int.coe_nat_sub na], + exact x_sub_y_dvd_pow a1 n k }, + have : n^k % t = m % t, from (this.symm.trans tm).add_left_cancel' _, + rw [← te] at nt, + rwa [nat.mod_eq_of_lt (nat.cast_lt.1 nt), nat.mod_eq_of_lt mt] at this } +end + +end pell From eb0cb4511aaef0da2462207b67358a0e1fe1e2ee Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sun, 26 Feb 2023 19:18:05 +0000 Subject: [PATCH 0539/1291] feat(matrix/dual_number): `matrix` and `dual_number` commute (#18386) This result is perhaps obvious, but it required a fair amount of prework to state. --- src/algebra/triv_sq_zero_ext.lean | 18 +++++++++++++-- src/data/matrix/dual_number.lean | 38 +++++++++++++++++++++++++++++++ 2 files changed, 54 insertions(+), 2 deletions(-) create mode 100644 src/data/matrix/dual_number.lean diff --git a/src/algebra/triv_sq_zero_ext.lean b/src/algebra/triv_sq_zero_ext.lean index c24d6f137553b..7cc0bf15019a2 100644 --- a/src/algebra/triv_sq_zero_ext.lean +++ b/src/algebra/triv_sq_zero_ext.lean @@ -60,6 +60,8 @@ R × M local notation `tsze` := triv_sq_zero_ext +open_locale big_operators + namespace triv_sq_zero_ext open mul_opposite (op) @@ -204,6 +206,12 @@ prod.module @[simp] lemma snd_smul [has_smul S R] [has_smul S M] (s : S) (x : tsze R M) : (s • x).snd = s • x.snd := rfl +lemma fst_sum {ι} [add_comm_monoid R] [add_comm_monoid M] (s : finset ι) (f : ι → tsze R M) : + (∑ i in s, f i).fst = ∑ i in s, (f i).fst := prod.fst_sum + +lemma snd_sum {ι} [add_comm_monoid R] [add_comm_monoid M] (s : finset ι) (f : ι → tsze R M) : + (∑ i in s, f i).snd = ∑ i in s, (f i).snd := prod.snd_sum + section variables (M) @@ -225,6 +233,10 @@ ext rfl (sub_zero _).symm (s : S) (r : R) : (inl (s • r) : tsze R M) = s • inl r := ext rfl (smul_zero s).symm +lemma inl_sum {ι} [add_comm_monoid R] [add_comm_monoid M] (s : finset ι) (f : ι → R) : + (inl (∑ i in s, f i) : tsze R M) = ∑ i in s, inl (f i) := +(linear_map.inl ℕ _ _).map_sum + end section @@ -248,6 +260,10 @@ ext (sub_zero _).symm rfl (r : S) (m : M) : (inr (r • m) : tsze R M) = r • inr m := ext (smul_zero _).symm rfl +lemma inr_sum {ι} [add_comm_monoid R] [add_comm_monoid M] (s : finset ι) (f : ι → M) : + (inr (∑ i in s, f i) : tsze R M) = ∑ i in s, inr (f i) := +(linear_map.inr ℕ _ _).map_sum + end lemma inl_fst_add_inr_snd_eq [add_zero_class R] [add_zero_class M] (x : tsze R M) : @@ -398,8 +414,6 @@ instance [ring R] [add_comm_group M] [module R M] [module Rᵐᵒᵖ M] : { .. triv_sq_zero_ext.add_group_with_one, .. triv_sq_zero_ext.non_assoc_semiring } -open_locale big_operators - /-- In the general non-commutative case, the power operator is $$\begin{align} diff --git a/src/data/matrix/dual_number.lean b/src/data/matrix/dual_number.lean new file mode 100644 index 0000000000000..256458b40e69e --- /dev/null +++ b/src/data/matrix/dual_number.lean @@ -0,0 +1,38 @@ +/- +Copyright (c) 2023 Eric Wieser. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Eric Wieser +-/ +import algebra.dual_number +import data.matrix.basic + +/-! +# Matrices of dual numbers are isomorphic to dual numbers over matrices + +Showing this for the more general case of `triv_sq_zero_ext R M` would require an action between +`matrix n n R` and `matrix n n M`, which would risk causing diamonds. +-/ + +variables {R n : Type} [comm_semiring R] [fintype n] [decidable_eq n] + +open matrix triv_sq_zero_ext + +/-- Matrices over dual numbers and dual numbers over matrices are isomorphic. -/ +@[simps] +def matrix.dual_number_equiv : matrix n n (dual_number R) ≃ₐ[R] dual_number (matrix n n R) := +{ to_fun := λ A, ⟨of (λ i j, (A i j).fst), of (λ i j, (A i j).snd)⟩, + inv_fun := λ d, of (λ i j, (d.fst i j, d.snd i j)), + left_inv := λ A, matrix.ext $ λ i j, triv_sq_zero_ext.ext rfl rfl, + right_inv := λ d, triv_sq_zero_ext.ext (matrix.ext $ λ i j, rfl) (matrix.ext $ λ i j, rfl), + map_mul' := λ A B, begin + ext; dsimp [mul_apply], + { simp_rw [fst_sum, fst_mul] }, + { simp_rw [snd_sum, snd_mul, smul_eq_mul, op_smul_eq_mul, finset.sum_add_distrib] }, + end, + map_add' := λ A B, triv_sq_zero_ext.ext rfl rfl, + commutes' := λ r, begin + simp_rw [algebra_map_eq_inl', algebra_map_eq_diagonal, pi.algebra_map_def, + algebra.id.map_eq_self, algebra_map_eq_inl, ←diagonal_map (inl_zero R), map_apply, + fst_inl, snd_inl], + refl, + end } From 0c1f285a9f6e608ae2bdffa3f993eafb01eba829 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sun, 26 Feb 2023 22:16:47 +0000 Subject: [PATCH 0540/1291] refactor(topology/metric_space): drop `dist_setoid` (#18502) Use `uniform_space.separation_setoid` instead. --- src/topology/constructions.lean | 6 +- src/topology/metric_space/basic.lean | 69 ++++------------ src/topology/metric_space/emetric_space.lean | 31 +++++++ src/topology/metric_space/gluing.lean | 38 ++++----- .../metric_space/gromov_hausdorff.lean | 31 +++---- .../gromov_hausdorff_realized.lean | 80 ++++++------------- src/topology/uniform_space/separation.lean | 10 +++ 7 files changed, 113 insertions(+), 152 deletions(-) diff --git a/src/topology/constructions.lean b/src/topology/constructions.lean index ad136bd9b80b4..3c8444fafec4e 100644 --- a/src/topology/constructions.lean +++ b/src/topology/constructions.lean @@ -757,9 +757,13 @@ lemma nhds_inl (x : α) : 𝓝 (inl x : α ⊕ β) = map inl (𝓝 x) := lemma nhds_inr (x : β) : 𝓝 (inr x : α ⊕ β) = map inr (𝓝 x) := (open_embedding_inr.map_nhds_eq _).symm +theorem continuous_sum_dom {f : α ⊕ β → γ} : + continuous f ↔ continuous (f ∘ sum.inl) ∧ continuous (f ∘ sum.inr) := +by simp only [continuous_sup_dom, continuous_coinduced_dom] + lemma continuous_sum_elim {f : α → γ} {g : β → γ} : continuous (sum.elim f g) ↔ continuous f ∧ continuous g := -by simp only [continuous_sup_dom, continuous_coinduced_dom, sum.elim_comp_inl, sum.elim_comp_inr] +continuous_sum_dom @[continuity] lemma continuous.sum_elim {f : α → γ} {g : β → γ} (hf : continuous f) (hg : continuous g) : continuous (sum.elim f g) := diff --git a/src/topology/metric_space/basic.lean b/src/topology/metric_space/basic.lean index b74e0400d871c..c5566c9b5ee0d 100644 --- a/src/topology/metric_space/basic.lean +++ b/src/topology/metric_space/basic.lean @@ -2962,61 +2962,20 @@ end metric section eq_rel -/-- The canonical equivalence relation on a pseudometric space. -/ -def pseudo_metric.dist_setoid (α : Type u) [pseudo_metric_space α] : setoid α := -setoid.mk (λx y, dist x y = 0) -begin - unfold equivalence, - repeat { split }, - { exact pseudo_metric_space.dist_self }, - { assume x y h, rwa pseudo_metric_space.dist_comm }, - { assume x y z hxy hyz, - refine le_antisymm _ dist_nonneg, - calc dist x z ≤ dist x y + dist y z : pseudo_metric_space.dist_triangle _ _ _ - ... = 0 + 0 : by rw [hxy, hyz] - ... = 0 : by simp } -end - -local attribute [instance] pseudo_metric.dist_setoid - -/-- The canonical quotient of a pseudometric space, identifying points at distance `0`. -/ -@[reducible] definition pseudo_metric_quot (α : Type u) [pseudo_metric_space α] : Type* := -quotient (pseudo_metric.dist_setoid α) - -instance has_dist_metric_quot {α : Type u} [pseudo_metric_space α] : - has_dist (pseudo_metric_quot α) := -{ dist := quotient.lift₂ (λp q : α, dist p q) -begin - assume x y x' y' hxx' hyy', - have Hxx' : dist x x' = 0 := hxx', - have Hyy' : dist y y' = 0 := hyy', - have A : dist x y ≤ dist x' y' := calc - dist x y ≤ dist x x' + dist x' y : pseudo_metric_space.dist_triangle _ _ _ - ... = dist x' y : by simp [Hxx'] - ... ≤ dist x' y' + dist y' y : pseudo_metric_space.dist_triangle _ _ _ - ... = dist x' y' : by simp [pseudo_metric_space.dist_comm, Hyy'], - have B : dist x' y' ≤ dist x y := calc - dist x' y' ≤ dist x' x + dist x y' : pseudo_metric_space.dist_triangle _ _ _ - ... = dist x y' : by simp [pseudo_metric_space.dist_comm, Hxx'] - ... ≤ dist x y + dist y y' : pseudo_metric_space.dist_triangle _ _ _ - ... = dist x y : by simp [Hyy'], - exact le_antisymm A B -end } - -lemma pseudo_metric_quot_dist_eq {α : Type u} [pseudo_metric_space α] (p q : α) : - dist ⟦p⟧ ⟦q⟧ = dist p q := rfl - -instance metric_space_quot {α : Type u} [pseudo_metric_space α] : - metric_space (pseudo_metric_quot α) := -{ dist_self := begin - refine quotient.ind (λy, _), - exact pseudo_metric_space.dist_self _ - end, - eq_of_dist_eq_zero := λxc yc, by exact quotient.induction_on₂ xc yc (λx y H, quotient.sound H), - dist_comm := - λxc yc, quotient.induction_on₂ xc yc (λx y, pseudo_metric_space.dist_comm _ _), - dist_triangle := - λxc yc zc, quotient.induction_on₃ xc yc zc (λx y z, pseudo_metric_space.dist_triangle _ _ _) } +instance {α : Type u} [pseudo_metric_space α] : + has_dist (uniform_space.separation_quotient α) := +{ dist := λ p q, quotient.lift_on₂' p q dist $ λ x y x' y' hx hy, + by rw [dist_edist, dist_edist, ← uniform_space.separation_quotient.edist_mk x, + ← uniform_space.separation_quotient.edist_mk x', quot.sound hx, quot.sound hy] } + +lemma uniform_space.separation_quotient.dist_mk {α : Type u} [pseudo_metric_space α] (p q : α) : + @dist (uniform_space.separation_quotient α) _ (quot.mk _ p) (quot.mk _ q) = dist p q := +rfl + +instance {α : Type u} [pseudo_metric_space α] : + metric_space (uniform_space.separation_quotient α) := +emetric_space.to_metric_space_of_dist dist (λ x y, quotient.induction_on₂' x y edist_ne_top) $ + λ x y, quotient.induction_on₂' x y dist_edist end eq_rel diff --git a/src/topology/metric_space/emetric_space.lean b/src/topology/metric_space/emetric_space.lean index 275a381e11557..acbfce8f93ca4 100644 --- a/src/topology/metric_space/emetric_space.lean +++ b/src/topology/metric_space/emetric_space.lean @@ -1062,6 +1062,37 @@ end diam end emetric +/-! +### Separation quotient +-/ + +instance [pseudo_emetric_space X] : has_edist (uniform_space.separation_quotient X) := +⟨λ x y, quotient.lift_on₂' x y edist $ λ x y x' y' hx hy, + calc edist x y = edist x' y : edist_congr_right $ + emetric.inseparable_iff.1 $ separation_rel_iff_inseparable.1 hx + ... = edist x' y' : edist_congr_left $ + emetric.inseparable_iff.1 $ separation_rel_iff_inseparable.1 hy⟩ + +@[simp] theorem uniform_space.separation_quotient.edist_mk [pseudo_emetric_space X] (x y : X) : + @edist (uniform_space.separation_quotient X) _ (quot.mk _ x) (quot.mk _ y) = edist x y := +rfl + +instance [pseudo_emetric_space X] : emetric_space (uniform_space.separation_quotient X) := +@emetric.of_t0_pseudo_emetric_space (uniform_space.separation_quotient X) + { edist_self := λ x, quotient.induction_on' x edist_self, + edist_comm := λ x y, quotient.induction_on₂' x y edist_comm, + edist_triangle := λ x y z, quotient.induction_on₃' x y z edist_triangle, + to_uniform_space := infer_instance, + uniformity_edist := (uniformity_basis_edist.map _).eq_binfi.trans $ infi_congr $ λ ε, + infi_congr $ λ hε, congr_arg 𝓟 + begin + ext ⟨⟨x⟩, ⟨y⟩⟩, + refine ⟨_, λ h, ⟨(x, y), h, rfl⟩⟩, + rintro ⟨⟨x', y'⟩, h', h⟩, + simp only [prod.ext_iff] at h, + rwa [← h.1, ← h.2] + end } _ + /-! ### `additive`, `multiplicative` diff --git a/src/topology/metric_space/gluing.lean b/src/topology/metric_space/gluing.lean index 6c628b30260c5..da4d50cd2a0dd 100644 --- a/src/topology/metric_space/gluing.lean +++ b/src/topology/metric_space/gluing.lean @@ -514,7 +514,7 @@ variables {X : Type u} {Y : Type v} {Z : Type w} variables [nonempty Z] [metric_space Z] [metric_space X] [metric_space Y] {Φ : Z → X} {Ψ : Z → Y} {ε : ℝ} open _root_.sum (inl inr) -local attribute [instance] pseudo_metric.dist_setoid +local attribute [instance] uniform_space.separation_setoid /-- Given two isometric embeddings `Φ : Z → X` and `Ψ : Z → Y`, we define a pseudo metric space structure on `X ⊕ Y` by declaring that `Φ x` and `Ψ x` are at distance `0`. -/ @@ -526,20 +526,15 @@ def glue_premetric (hΦ : isometry Φ) (hΨ : isometry Ψ) : pseudo_metric_space /-- Given two isometric embeddings `Φ : Z → X` and `Ψ : Z → Y`, we define a space `glue_space hΦ hΨ` by identifying in `X ⊕ Y` the points `Φ x` and `Ψ x`. -/ +@[derive metric_space] def glue_space (hΦ : isometry Φ) (hΨ : isometry Ψ) : Type* := -@pseudo_metric_quot _ (glue_premetric hΦ hΨ) - -instance metric_space_glue_space (hΦ : isometry Φ) (hΨ : isometry Ψ) : - metric_space (glue_space hΦ hΨ) := -@metric_space_quot _ (glue_premetric hΦ hΨ) +@uniform_space.separation_quotient _ (glue_premetric hΦ hΨ).to_uniform_space /-- The canonical map from `X` to the space obtained by gluing isometric subsets in `X` and `Y`. -/ -def to_glue_l (hΦ : isometry Φ) (hΨ : isometry Ψ) (x : X) : glue_space hΦ hΨ := -by letI : pseudo_metric_space (X ⊕ Y) := glue_premetric hΦ hΨ; exact ⟦inl x⟧ +def to_glue_l (hΦ : isometry Φ) (hΨ : isometry Ψ) (x : X) : glue_space hΦ hΨ := quotient.mk' (inl x) /-- The canonical map from `Y` to the space obtained by gluing isometric subsets in `X` and `Y`. -/ -def to_glue_r (hΦ : isometry Φ) (hΨ : isometry Ψ) (y : Y) : glue_space hΦ hΨ := -by letI : pseudo_metric_space (X ⊕ Y) := glue_premetric hΦ hΨ; exact ⟦inr y⟧ +def to_glue_r (hΦ : isometry Φ) (hΨ : isometry Ψ) (y : Y) : glue_space hΦ hΨ := quotient.mk' (inr y) instance inhabited_left (hΦ : isometry Φ) (hΨ : isometry Ψ) [inhabited X] : inhabited (glue_space hΦ hΨ) := @@ -552,9 +547,11 @@ instance inhabited_right (hΦ : isometry Φ) (hΨ : isometry Ψ) [inhabited Y] : lemma to_glue_commute (hΦ : isometry Φ) (hΨ : isometry Ψ) : (to_glue_l hΦ hΨ) ∘ Φ = (to_glue_r hΦ hΨ) ∘ Ψ := begin - letI : pseudo_metric_space (X ⊕ Y) := glue_premetric hΦ hΨ, + letI i : pseudo_metric_space (X ⊕ Y) := glue_premetric hΦ hΨ, + letI := i.to_uniform_space, funext, - simp only [comp, to_glue_l, to_glue_r, quotient.eq], + simp only [comp, to_glue_l, to_glue_r], + refine uniform_space.separation_quotient.mk_eq_mk.2 (metric.inseparable_iff.2 _), exact glue_dist_glued_points Φ Ψ 0 x end @@ -636,20 +633,15 @@ def inductive_premetric (I : ∀ n, isometry (f n)) : inductive_limit_dist_eq_dist I y z m hy hz] end } -local attribute [instance] inductive_premetric pseudo_metric.dist_setoid +local attribute [instance] inductive_premetric uniform_space.separation_setoid /-- The type giving the inductive limit in a metric space context. -/ -def inductive_limit (I : ∀ n, isometry (f n)) : Type* := -@pseudo_metric_quot _ (inductive_premetric I) - -/-- Metric space structure on the inductive limit. -/ -instance metric_space_inductive_limit (I : ∀ n, isometry (f n)) : - metric_space (inductive_limit I) := -@metric_space_quot _ (inductive_premetric I) +@[derive metric_space] def inductive_limit (I : ∀ n, isometry (f n)) : Type* := +@uniform_space.separation_quotient _ (inductive_premetric I).to_uniform_space /-- Mapping each `X n` to the inductive limit. -/ def to_inductive_limit (I : ∀ n, isometry (f n)) (n : ℕ) (x : X n) : metric.inductive_limit I := -by letI : pseudo_metric_space (Σ n, X n) := inductive_premetric I; exact ⟦sigma.mk n x⟧ +quotient.mk' (sigma.mk n x) instance (I : ∀ n, isometry (f n)) [inhabited (X 0)] : inhabited (inductive_limit I) := ⟨to_inductive_limit _ 0 default⟩ @@ -667,8 +659,10 @@ end lemma to_inductive_limit_commute (I : ∀ n, isometry (f n)) (n : ℕ) : (to_inductive_limit I n.succ) ∘ (f n) = to_inductive_limit I n := begin + letI := inductive_premetric I, funext, - simp only [comp, to_inductive_limit, quotient.eq], + simp only [comp, to_inductive_limit], + refine uniform_space.separation_quotient.mk_eq_mk.2 (metric.inseparable_iff.2 _), show inductive_limit_dist f ⟨n.succ, f n x⟩ ⟨n, x⟩ = 0, { rw [inductive_limit_dist_eq_dist I ⟨n.succ, f n x⟩ ⟨n, x⟩ n.succ, le_rec_on_self, le_rec_on_succ, le_rec_on_self, dist_self], diff --git a/src/topology/metric_space/gromov_hausdorff.lean b/src/topology/metric_space/gromov_hausdorff.lean index ace9c9a3d64e9..d6b068acf48f9 100644 --- a/src/topology/metric_space/gromov_hausdorff.lean +++ b/src/topology/metric_space/gromov_hausdorff.lean @@ -429,7 +429,6 @@ instance : metric_space GH_space := let Ψ : Y → γ2 := optimal_GH_injl Y Z, have hΨ : isometry Ψ := isometry_optimal_GH_injl Y Z, let γ := glue_space hΦ hΨ, - letI : metric_space γ := metric.metric_space_glue_space hΦ hΨ, have Comm : (to_glue_l hΦ hΨ) ∘ (optimal_GH_injr X Y) = (to_glue_r hΦ hΨ) ∘ (optimal_GH_injl Y Z) := to_glue_commute hΦ hΨ, calc dist x z = dist (to_GH_space X) (to_GH_space Z) : @@ -918,6 +917,8 @@ structure aux_gluing_struct (A : Type) [metric_space A] : Type 1 := (embed : A → space) (isom : isometry embed) +local attribute [instance] aux_gluing_struct.metric + instance (A : Type) [metric_space A] : inhabited (aux_gluing_struct A) := ⟨{ space := A, metric := by apply_instance, @@ -927,17 +928,13 @@ instance (A : Type) [metric_space A] : inhabited (aux_gluing_struct A) := /-- Auxiliary sequence of metric spaces, containing copies of `X 0`, ..., `X n`, where each `X i` is glued to `X (i+1)` in an optimal way. The space at step `n+1` is obtained from the space at step `n` by adding `X (n+1)`, glued in an optimal way to the `X n` already sitting there. -/ -def aux_gluing (n : ℕ) : aux_gluing_struct (X n) := nat.rec_on n - { space := X 0, - metric := by apply_instance, - embed := id, - isom := λ x y, rfl } -(λ n Y, by letI : metric_space Y.space := Y.metric; exact +def aux_gluing (n : ℕ) : aux_gluing_struct (X n) := +nat.rec_on n default $ λ n Y, { space := glue_space Y.isom (isometry_optimal_GH_injl (X n) (X (n+1))), metric := by apply_instance, embed := (to_glue_r Y.isom (isometry_optimal_GH_injl (X n) (X (n+1)))) ∘ (optimal_GH_injr (X n) (X (n+1))), - isom := (to_glue_r_isometry _ _).comp (isometry_optimal_GH_injr (X n) (X (n+1))) }) + isom := (to_glue_r_isometry _ _).comp (isometry_optimal_GH_injr (X n) (X (n+1))) } /-- The Gromov-Hausdorff space is complete. -/ instance : complete_space GH_space := @@ -949,20 +946,16 @@ begin let X := λ n, (u n).rep, -- glue them together successively in an optimal way, getting a sequence of metric spaces `Y n` let Y := aux_gluing X, - letI : ∀ n, metric_space (Y n).space := λ n, (Y n).metric, + -- this equality is true by definition but Lean unfolds some defs in the wrong order have E : ∀ n : ℕ, - glue_space (Y n).isom (isometry_optimal_GH_injl (X n) (X n.succ)) = (Y n.succ).space := - λ n, by { simp only [Y, aux_gluing], refl }, + glue_space (Y n).isom (isometry_optimal_GH_injl (X n) (X (n + 1))) = (Y (n + 1)).space := + λ n, by { dsimp only [Y, aux_gluing], refl }, let c := λ n, cast (E n), - have ic : ∀ n, isometry (c n) := λ n x y, rfl, + have ic : ∀ n, isometry (c n) := λ n x y, by { dsimp only [Y, aux_gluing], exact rfl }, -- there is a canonical embedding of `Y n` in `Y (n+1)`, by construction - let f : Πn, (Y n).space → (Y n.succ).space := - λ n, (c n) ∘ (to_glue_l (aux_gluing X n).isom (isometry_optimal_GH_injl (X n) (X n.succ))), - have I : ∀ n, isometry (f n), - { assume n, - apply isometry.comp, - { assume x y, refl }, - { apply to_glue_l_isometry } }, + let f : Π n, (Y n).space → (Y (n + 1)).space := + λ n, c n ∘ to_glue_l (Y n).isom (isometry_optimal_GH_injl (X n) (X n.succ)), + have I : ∀ n, isometry (f n) := λ n, (ic n).comp (to_glue_l_isometry _ _), -- consider the inductive limit `Z0` of the `Y n`, and then its completion `Z` let Z0 := metric.inductive_limit I, let Z := uniform_space.completion Z0, diff --git a/src/topology/metric_space/gromov_hausdorff_realized.lean b/src/topology/metric_space/gromov_hausdorff_realized.lean index f53cd689243bf..0a25b1becc138 100644 --- a/src/topology/metric_space/gromov_hausdorff_realized.lean +++ b/src/topology/metric_space/gromov_hausdorff_realized.lean @@ -90,14 +90,14 @@ local attribute [instance, priority 10] inhabited_of_nonempty' private lemma max_var_bound : dist x y ≤ max_var X Y := calc dist x y ≤ diam (univ : set (X ⊕ Y)) : dist_le_diam_of_mem bounded_of_compact_space (mem_univ _) (mem_univ _) - ... = diam (inl '' (univ : set X) ∪ inr '' (univ : set Y)) : - by apply congr_arg; ext x y z; cases x; simp [mem_univ, mem_range_self] - ... ≤ diam (inl '' (univ : set X)) + dist (inl default) (inr default) + - diam (inr '' (univ : set Y)) : - diam_union (mem_image_of_mem _ (mem_univ _)) (mem_image_of_mem _ (mem_univ _)) + ... = diam (range inl ∪ range inr : set (X ⊕ Y)) : + by rw [range_inl_union_range_inr] + ... ≤ diam (range inl : set (X ⊕ Y)) + dist (inl default) (inr default) + + diam (range inr : set (X ⊕ Y)) : + diam_union (mem_range_self _) (mem_range_self _) ... = diam (univ : set X) + (dist default default + 1 + dist default default) + diam (univ : set Y) : - by { rw [isometry_inl.diam_image, isometry_inr.diam_image], refl } + by { rw [isometry_inl.diam_range, isometry_inr.diam_range], refl } ... = 1 * diam (univ : set X) + 1 + 1 * diam (univ : set Y) : by simp ... ≤ 2 * diam (univ : set X) + 1 + 2 * diam (univ : set Y) : begin @@ -448,52 +448,33 @@ def premetric_optimal_GH_dist : pseudo_metric_space (X ⊕ Y) := dist_comm := λx y, candidates_symm (optimal_GH_dist_mem_candidates_b X Y), dist_triangle := λx y z, candidates_triangle (optimal_GH_dist_mem_candidates_b X Y) } -local attribute [instance] premetric_optimal_GH_dist pseudo_metric.dist_setoid +local attribute [instance] premetric_optimal_GH_dist /-- A metric space which realizes the optimal coupling between `X` and `Y` -/ @[derive metric_space, nolint has_nonempty_instance] definition optimal_GH_coupling : Type* := -pseudo_metric_quot (X ⊕ Y) +@uniform_space.separation_quotient (X ⊕ Y) (premetric_optimal_GH_dist X Y).to_uniform_space /-- Injection of `X` in the optimal coupling between `X` and `Y` -/ -def optimal_GH_injl (x : X) : optimal_GH_coupling X Y := ⟦inl x⟧ +def optimal_GH_injl (x : X) : optimal_GH_coupling X Y := quotient.mk' (inl x) /-- The injection of `X` in the optimal coupling between `X` and `Y` is an isometry. -/ lemma isometry_optimal_GH_injl : isometry (optimal_GH_injl X Y) := -begin - refine isometry.of_dist_eq (λx y, _), - change dist ⟦inl x⟧ ⟦inl y⟧ = dist x y, - exact candidates_dist_inl (optimal_GH_dist_mem_candidates_b X Y) _ _, -end +isometry.of_dist_eq $ λ x y, candidates_dist_inl (optimal_GH_dist_mem_candidates_b X Y) _ _ /-- Injection of `Y` in the optimal coupling between `X` and `Y` -/ -def optimal_GH_injr (y : Y) : optimal_GH_coupling X Y := ⟦inr y⟧ +def optimal_GH_injr (y : Y) : optimal_GH_coupling X Y := quotient.mk' (inr y) /-- The injection of `Y` in the optimal coupling between `X` and `Y` is an isometry. -/ lemma isometry_optimal_GH_injr : isometry (optimal_GH_injr X Y) := -begin - refine isometry.of_dist_eq (λx y, _), - change dist ⟦inr x⟧ ⟦inr y⟧ = dist x y, - exact candidates_dist_inr (optimal_GH_dist_mem_candidates_b X Y) _ _, -end +isometry.of_dist_eq $ λ x y, candidates_dist_inr (optimal_GH_dist_mem_candidates_b X Y) _ _ /-- The optimal coupling between two compact spaces `X` and `Y` is still a compact space -/ instance compact_space_optimal_GH_coupling : compact_space (optimal_GH_coupling X Y) := ⟨begin - have : (univ : set (optimal_GH_coupling X Y)) = - (optimal_GH_injl X Y '' univ) ∪ (optimal_GH_injr X Y '' univ), - { refine subset.antisymm (λxc hxc, _) (subset_univ _), - rcases quotient.exists_rep xc with ⟨x, hx⟩, - cases x; rw ← hx, - { have : ⟦inl x⟧ = optimal_GH_injl X Y x := rfl, - rw this, - exact mem_union_left _ (mem_image_of_mem _ (mem_univ _)) }, - { have : ⟦inr x⟧ = optimal_GH_injr X Y x := rfl, - rw this, - exact mem_union_right _ (mem_image_of_mem _ (mem_univ _)) } }, - rw this, - exact (is_compact_univ.image (isometry_optimal_GH_injl X Y).continuous).union - (is_compact_univ.image (isometry_optimal_GH_injr X Y).continuous) + rw [← range_quotient_mk'], + exact is_compact_range (continuous_sum_dom.2 ⟨(isometry_optimal_GH_injl X Y).continuous, + (isometry_optimal_GH_injr X Y).continuous⟩) end⟩ /-- For any candidate `f`, `HD(f)` is larger than or equal to the Hausdorff distance in the @@ -503,42 +484,31 @@ we need. -/ lemma Hausdorff_dist_optimal_le_HD {f} (h : f ∈ candidates_b X Y) : Hausdorff_dist (range (optimal_GH_injl X Y)) (range (optimal_GH_injr X Y)) ≤ HD f := begin - refine le_trans (le_of_forall_le_of_dense (λr hr, _)) (HD_optimal_GH_dist_le X Y f h), + refine le_trans (le_of_forall_le_of_dense (λ r hr, _)) (HD_optimal_GH_dist_le X Y f h), have A : ∀ x ∈ range (optimal_GH_injl X Y), ∃ y ∈ range (optimal_GH_injr X Y), dist x y ≤ r, - { assume x hx, - rcases mem_range.1 hx with ⟨z, hz⟩, - rw ← hz, + { rintro _ ⟨z, rfl⟩, have I1 : (⨆ x, ⨅ y, optimal_GH_dist X Y (inl x, inr y)) < r := lt_of_le_of_lt (le_max_left _ _) hr, have I2 : (⨅ y, optimal_GH_dist X Y (inl z, inr y)) ≤ ⨆ x, ⨅ y, optimal_GH_dist X Y (inl x, inr y) := le_cSup (by simpa using HD_bound_aux1 _ 0) (mem_range_self _), have I : (⨅ y, optimal_GH_dist X Y (inl z, inr y)) < r := lt_of_le_of_lt I2 I1, - rcases exists_lt_of_cInf_lt (range_nonempty _) I with ⟨r', r'range, hr'⟩, - rcases mem_range.1 r'range with ⟨z', hz'⟩, - existsi [optimal_GH_injr X Y z', mem_range_self _], - have : (optimal_GH_dist X Y) (inl z, inr z') ≤ r, by { rw hz', exact le_of_lt hr' }, - exact this }, + rcases exists_lt_of_cInf_lt (range_nonempty _) I with ⟨r', ⟨z', rfl⟩, hr'⟩, + exact ⟨optimal_GH_injr X Y z', mem_range_self _, le_of_lt hr'⟩ }, refine Hausdorff_dist_le_of_mem_dist _ A _, - { rcases exists_mem_of_nonempty X with ⟨xX, _⟩, - have : optimal_GH_injl X Y xX ∈ range (optimal_GH_injl X Y) := mem_range_self _, - rcases A _ this with ⟨y, yrange, hy⟩, + { inhabit X, + rcases A _ (mem_range_self default) with ⟨y, -, hy⟩, exact le_trans dist_nonneg hy }, - { assume y hy, - rcases mem_range.1 hy with ⟨z, hz⟩, - rw ← hz, + { rintro _ ⟨z, rfl⟩, have I1 : (⨆ y, ⨅ x, optimal_GH_dist X Y (inl x, inr y)) < r := lt_of_le_of_lt (le_max_right _ _) hr, have I2 : (⨅ x, optimal_GH_dist X Y (inl x, inr z)) ≤ ⨆ y, ⨅ x, optimal_GH_dist X Y (inl x, inr y) := le_cSup (by simpa using HD_bound_aux2 _ 0) (mem_range_self _), have I : (⨅ x, optimal_GH_dist X Y (inl x, inr z)) < r := lt_of_le_of_lt I2 I1, - rcases exists_lt_of_cInf_lt (range_nonempty _) I with ⟨r', r'range, hr'⟩, - rcases mem_range.1 r'range with ⟨z', hz'⟩, - existsi [optimal_GH_injl X Y z', mem_range_self _], - have : (optimal_GH_dist X Y) (inl z', inr z) ≤ r, by { rw hz', exact le_of_lt hr' }, - rw dist_comm, - exact this } + rcases exists_lt_of_cInf_lt (range_nonempty _) I with ⟨r', ⟨z', rfl⟩, hr'⟩, + refine ⟨optimal_GH_injl X Y z', mem_range_self _, le_of_lt _⟩, + rwa dist_comm } end end consequences diff --git a/src/topology/uniform_space/separation.lean b/src/topology/uniform_space/separation.lean index 54c7cfe6ded99..32fda3d944457 100644 --- a/src/topology/uniform_space/separation.lean +++ b/src/topology/uniform_space/separation.lean @@ -113,6 +113,13 @@ lemma filter.has_basis.mem_separation_rel {ι : Sort*} {p : ι → Prop} {s : ι a ∈ 𝓢 α ↔ ∀ i, p i → a ∈ s i := h.forall_mem_mem +theorem separation_rel_iff_specializes {a b : α} : (a, b) ∈ 𝓢 α ↔ a ⤳ b := +by simp only [(𝓤 α).basis_sets.mem_separation_rel, id, mem_set_of_eq, + (nhds_basis_uniformity (𝓤 α).basis_sets).specializes_iff] + +theorem separation_rel_iff_inseparable {a b : α} : (a, b) ∈ 𝓢 α ↔ inseparable a b := + separation_rel_iff_specializes.trans specializes_iff_inseparable + /-- A uniform space is separated if its separation relation is trivial (each point is related only to itself). -/ class separated_space (α : Type u) [uniform_space α] : Prop := (out : 𝓢 α = id_rel) @@ -362,6 +369,9 @@ instance : separated_space (separation_quotient α) := uniform_space.separated_s instance [inhabited α] : inhabited (separation_quotient α) := quotient.inhabited (separation_setoid α) +lemma mk_eq_mk {x y : α} : (⟦x⟧ : separation_quotient α) = ⟦y⟧ ↔ inseparable x y := +quotient.eq'.trans separation_rel_iff_inseparable + /-- Factoring functions to a separated space through the separation quotient. -/ def lift [separated_space β] (f : α → β) : (separation_quotient α → β) := if h : uniform_continuous f then From ea94d7cd54ad9ca6b7710032868abb7c6a104c9c Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 27 Feb 2023 00:20:09 +0000 Subject: [PATCH 0541/1291] docs(algebra/module/equiv): correct a misleading comment (#18458) The linter is saying that the argument is unused because the argument is unused. The comment claiming that it doesn't appear and the linter is just confused is false. We could remove the argument, but the extra generality it would provide is illusory, and it would likely just be inconvenient. This is forward-ported in https://github.com/leanprover-community/mathlib4/pull/2347, though we will need to update the SHA again after this PR is merged. --- src/algebra/module/equiv.lean | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/src/algebra/module/equiv.lean b/src/algebra/module/equiv.lean index 636be037e8be8..5cb99e06014d4 100644 --- a/src/algebra/module/equiv.lean +++ b/src/algebra/module/equiv.lean @@ -240,7 +240,8 @@ variables (e₁₂ : M₁ ≃ₛₗ[σ₁₂] M₂) (e₂₃ : M₂ ≃ₛₗ[σ include σ₃₁ /-- Linear equivalences are transitive. -/ --- Note: The linter thinks the `ring_hom_comp_triple` argument is doubled -- it is not. +-- Note: the `ring_hom_comp_triple σ₃₂ σ₂₁ σ₃₁` is unused, but is convenient to carry around +-- implicitly for lemmas like `linear_equiv.self_trans_symm`. @[trans, nolint unused_arguments] def trans : M₁ ≃ₛₗ[σ₁₃] M₃ := { .. e₂₃.to_linear_map.comp e₁₂.to_linear_map, From 45a1ada01ed893722d0f8d15b9e44270996515d5 Mon Sep 17 00:00:00 2001 From: Thomas Browning Date: Mon, 27 Feb 2023 02:57:37 +0000 Subject: [PATCH 0542/1291] feat(data/int/units): If `z * w = 1`, then `z = w` (#18499) This PR adds a lemma stating that if `z * w = 1`, then `z = w`. mathlib4 PR: https://github.com/leanprover-community/mathlib4/pull/2497 --- src/data/int/units.lean | 3 +++ 1 file changed, 3 insertions(+) diff --git a/src/data/int/units.lean b/src/data/int/units.lean index 7184fa0163efd..741d16d48802f 100644 --- a/src/data/int/units.lean +++ b/src/data/int/units.lean @@ -57,6 +57,9 @@ begin tauto, end +theorem eq_of_mul_eq_one {z w : ℤ} (h : z * w = 1) : z = w := +(eq_one_or_neg_one_of_mul_eq_one' h).elim (λ h, h.1.trans h.2.symm) (λ h, h.1.trans h.2.symm) + lemma mul_eq_one_iff_eq_one_or_neg_one {z w : ℤ} : z * w = 1 ↔ z = 1 ∧ w = 1 ∨ z = -1 ∧ w = -1 := begin From 03994e05d8bfc59a41d2ec99523d6553d21848ac Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Mon, 27 Feb 2023 08:41:38 +0000 Subject: [PATCH 0543/1291] feat(algebraic_geometry/elliptic_curve/point): define homomorphism induced by base change (#18493) --- .../elliptic_curve/point.lean | 167 ++++++++++++++++-- .../elliptic_curve/weierstrass.lean | 41 ++++- 2 files changed, 191 insertions(+), 17 deletions(-) diff --git a/src/algebraic_geometry/elliptic_curve/point.lean b/src/algebraic_geometry/elliptic_curve/point.lean index 7079473cef4fd..665ff960c8798 100644 --- a/src/algebraic_geometry/elliptic_curve/point.lean +++ b/src/algebraic_geometry/elliptic_curve/point.lean @@ -67,6 +67,9 @@ The group law on this set is then uniquely determined by these constructions. elliptic curve, rational point, group law -/ +private meta def map_simp : tactic unit := +`[simp only [map_one, map_bit0, map_bit1, map_neg, map_add, map_sub, map_mul, map_pow, map_div₀]] + private meta def eval_simp : tactic unit := `[simp only [eval_C, eval_X, eval_neg, eval_add, eval_sub, eval_mul, eval_pow]] @@ -77,7 +80,7 @@ private meta def derivative_simp : tactic unit := `[simp only [derivative_C, derivative_X, derivative_X_pow, derivative_neg, derivative_add, derivative_sub, derivative_mul, derivative_sq]] -universe u +universes u v w namespace weierstrass_curve @@ -89,7 +92,9 @@ section basic /-! ### Polynomials associated to nonsingular rational points on a Weierstrass curve -/ -variables {R : Type u} [comm_ring R] (W : weierstrass_curve R) (x₁ x₂ y₁ y₂ L : R) +variables {R : Type u} [comm_ring R] (W : weierstrass_curve R) (A : Type v) [comm_ring A] + [algebra R A] (B : Type w) [comm_ring B] [algebra R B] [algebra A B] [is_scalar_tower R A B] + (x₁ x₂ y₁ y₂ L : R) /-- The polynomial $-Y - a_1X - a_3$ associated to negation. -/ noncomputable def neg_polynomial : R[X][Y] := -Y - C (C W.a₁ * X + C W.a₃) @@ -101,16 +106,26 @@ This depends on `W`, and has argument order: $x_1$, $y_1$. -/ lemma neg_Y_neg_Y : W.neg_Y x₁ (W.neg_Y x₁ y₁) = y₁ := by { simp only [neg_Y], ring1 } +lemma base_change_neg_Y : + (W.base_change A).neg_Y (algebra_map R A x₁) (algebra_map R A y₁) + = algebra_map R A (W.neg_Y x₁ y₁) := +by { simp only [neg_Y], map_simp, refl } + +lemma base_change_neg_Y_of_base_change (x₁ y₁ : A) : + (W.base_change B).neg_Y (algebra_map A B x₁) (algebra_map A B y₁) + = algebra_map A B ((W.base_change A).neg_Y x₁ y₁) := +by rw [← base_change_neg_Y, base_change_base_change] + @[simp] lemma eval_neg_polynomial : (W.neg_polynomial.eval $ C y₁).eval x₁ = W.neg_Y x₁ y₁ := by { rw [neg_Y, sub_sub, neg_polynomial], eval_simp } -/-- The polynomial $L*(X - x_1) + y_1$ associated to the line $Y = L*(X - x_1) + y_1$, +/-- The polynomial $L(X - x_1) + y_1$ associated to the line $Y = L(X - x_1) + y_1$, with a slope of $L$ that passes through an affine point $(x_1, y_1)$. This does not depend on `W`, and has argument order: $x_1$, $y_1$, $L$. -/ noncomputable def line_polynomial : R[X] := C L * (X - C x₁) + C y₁ -/-- The polynomial obtained by substituting the line $Y = L*(X - x_1) + y_1$, with a slope of $L$ +/-- The polynomial obtained by substituting the line $Y = L(X - x_1) + y_1$, with a slope of $L$ that passes through an affine point $(x_1, y_1)$, into the polynomial $W(X, Y)$ associated to `W`. If such a line intersects `W` at another point $(x_2, y_2)$, then the roots of this polynomial are precisely $x_1$, $x_2$, and the $X$-coordinate of the addition of $(x_1, y_1)$ and $(x_2, y_2)$. @@ -131,18 +146,48 @@ where the line through them is not vertical and has a slope of $L$. This depends on `W`, and has argument order: $x_1$, $x_2$, $L$. -/ @[simp] def add_X : R := L ^ 2 + W.a₁ * L - W.a₂ - x₁ - x₂ +lemma base_change_add_X : + (W.base_change A).add_X (algebra_map R A x₁) (algebra_map R A x₂) (algebra_map R A L) + = algebra_map R A (W.add_X x₁ x₂ L) := +by { simp only [add_X], map_simp, refl } + +lemma base_change_add_X_of_base_change (x₁ x₂ L : A) : + (W.base_change B).add_X (algebra_map A B x₁) (algebra_map A B x₂) (algebra_map A B L) + = algebra_map A B ((W.base_change A).add_X x₁ x₂ L) := +by rw [← base_change_add_X, base_change_base_change] + /-- The $Y$-coordinate, before applying the final negation, of the addition of two affine points $(x_1, y_1)$ and $(x_2, y_2)$, where the line through them is not vertical and has a slope of $L$. This depends on `W`, and has argument order: $x_1$, $x_2$, $y_1$, $L$. -/ @[simp] def add_Y' : R := L * (W.add_X x₁ x₂ L - x₁) + y₁ +lemma base_change_add_Y' : + (W.base_change A).add_Y' (algebra_map R A x₁) (algebra_map R A x₂) (algebra_map R A y₁) + (algebra_map R A L) = algebra_map R A (W.add_Y' x₁ x₂ y₁ L) := +by { simp only [add_Y', base_change_add_X], map_simp } + +lemma base_change_add_Y'_of_base_change (x₁ x₂ y₁ L : A) : + (W.base_change B).add_Y' (algebra_map A B x₁) (algebra_map A B x₂) (algebra_map A B y₁) + (algebra_map A B L) = algebra_map A B ((W.base_change A).add_Y' x₁ x₂ y₁ L) := +by rw [← base_change_add_Y', base_change_base_change] + /-- The $Y$-coordinate of the addition of two affine points $(x_1, y_1)$ and $(x_2, y_2)$ in `W`, where the line through them is not vertical and has a slope of $L$. This depends on `W`, and has argument order: $x_1$, $x_2$, $y_1$, $L$. -/ @[simp] def add_Y : R := W.neg_Y (W.add_X x₁ x₂ L) (W.add_Y' x₁ x₂ y₁ L) +lemma base_change_add_Y : + (W.base_change A).add_Y (algebra_map R A x₁) (algebra_map R A x₂) (algebra_map R A y₁) + (algebra_map R A L) = algebra_map R A (W.add_Y x₁ x₂ y₁ L) := +by simp only [add_Y, base_change_add_Y', base_change_add_X, base_change_neg_Y] + +lemma base_change_add_Y_of_base_change (x₁ x₂ y₁ L : A) : + (W.base_change B).add_Y (algebra_map A B x₁) (algebra_map A B x₂) (algebra_map A B y₁) + (algebra_map A B L) = algebra_map A B ((W.base_change A).add_Y x₁ x₂ y₁ L) := +by rw [← base_change_add_Y, base_change_base_change] + lemma equation_add_iff : W.equation (W.add_X x₁ x₂ L) (W.add_Y' x₁ x₂ y₁ L) ↔ (W.add_polynomial x₁ y₁ L).eval (W.add_X x₁ x₂ L) = 0 := @@ -241,7 +286,8 @@ section addition open_locale classical -variables {F : Type u} [field F] (W : weierstrass_curve F) (x₁ x₂ y₁ y₂ : F) +variables {F : Type u} [field F] (W : weierstrass_curve F) (K : Type v) [field K] [algebra F K] + (x₁ x₂ y₁ y₂ : F) /-- The slope of the line through two affine points $(x_1, y_1)$ and $(x_2, y_2)$ in `W`. If $x_1 \ne x_2$, then this line is the secant of `W` through $(x_1, y_1)$ and $(x_2, y_2)$, @@ -259,6 +305,10 @@ else (y₁ - y₂) / (x₁ - x₂) variables {W x₁ x₂ y₁ y₂} (h₁ : W.nonsingular x₁ y₁) (h₂ : W.nonsingular x₂ y₂) (h₁' : W.equation x₁ y₁) (h₂' : W.equation x₂ y₂) +@[simp] lemma slope_of_Y_eq (hx : x₁ = x₂) (hy : y₁ = W.neg_Y x₂ y₂) : + W.slope x₁ x₂ y₁ y₂ = 0 := +by rw [slope, dif_pos hx, dif_pos hy] + @[simp] lemma slope_of_Y_ne (hx : x₁ = x₂) (hy : y₁ ≠ W.neg_Y x₂ y₂) : W.slope x₁ x₂ y₁ y₂ = (3 * x₁ ^ 2 + 2 * W.a₂ * x₁ + W.a₄ - W.a₁ * y₁) / (y₁ - W.neg_Y x₁ y₁) := by rw [slope, dif_pos hx, dif_neg hy] @@ -272,6 +322,33 @@ lemma slope_of_Y_ne_eq_eval (hx : x₁ = x₂) (hy : y₁ ≠ W.neg_Y x₂ y₂) by { rw [slope_of_Y_ne hx hy, eval_polynomial_X, neg_sub], congr' 1, rw [neg_Y, eval_polynomial_Y], ring1 } +lemma base_change_slope : + (W.base_change K).slope (algebra_map F K x₁) (algebra_map F K x₂) (algebra_map F K y₁) + (algebra_map F K y₂) = algebra_map F K (W.slope x₁ x₂ y₁ y₂) := +begin + by_cases hx : x₁ = x₂, + { by_cases hy : y₁ = W.neg_Y x₂ y₂, + { rw [slope_of_Y_eq hx hy, slope_of_Y_eq $ congr_arg _ hx, map_zero], + { rw [hy, base_change_neg_Y] } }, + { rw [slope_of_Y_ne hx hy, slope_of_Y_ne $ congr_arg _ hx], + { map_simp, + simpa only [base_change_neg_Y] }, + { rw [base_change_neg_Y], + contrapose! hy, + exact no_zero_smul_divisors.algebra_map_injective F K hy } } }, + { rw [slope_of_X_ne hx, slope_of_X_ne], + { map_simp }, + { contrapose! hx, + exact no_zero_smul_divisors.algebra_map_injective F K hx } } +end + +lemma base_change_slope_of_base_change {R : Type u} [comm_ring R] (W : weierstrass_curve R) + (F : Type v) [field F] [algebra R F] (K : Type w) [field K] [algebra R K] [algebra F K] + [is_scalar_tower R F K] (x₁ x₂ y₁ y₂ : F) : + (W.base_change K).slope (algebra_map F K x₁) (algebra_map F K x₂) (algebra_map F K y₁) + (algebra_map F K y₂) = algebra_map F K ((W.base_change F).slope x₁ x₂ y₁ y₂) := +by rw [← base_change_slope, base_change_base_change] + include h₁' h₂' lemma Y_eq_of_X_eq (hx : x₁ = x₂) : y₁ = y₂ ∨ y₁ = W.neg_Y x₂ y₂ := @@ -368,6 +445,8 @@ omit h₁ h₂ namespace point +variables {h₁ h₂} + /-- The addition of two nonsingular rational points. Given two nonsingular rational points `P` and `Q`, use `P + Q` instead of `add P Q`. -/ @@ -391,7 +470,7 @@ noncomputable instance : add_zero_class W.point := by rw [← add_def, add, dif_pos hx, dif_pos hy] @[simp] lemma some_add_self_of_Y_eq (hy : y₁ = W.neg_Y x₁ y₁) : some h₁ + some h₁ = 0 := -some_add_some_of_Y_eq h₁ h₁ rfl hy +some_add_some_of_Y_eq rfl hy @[simp] lemma some_add_some_of_Y_ne (hx : x₁ = x₂) (hy : y₁ ≠ W.neg_Y x₂ y₂) : some h₁ + some h₂ = some (nonsingular_add h₁ h₂ $ λ _, hy) := @@ -399,15 +478,15 @@ by rw [← add_def, add, dif_pos hx, dif_neg hy] lemma some_add_some_of_Y_ne' (hx : x₁ = x₂) (hy : y₁ ≠ W.neg_Y x₂ y₂) : some h₁ + some h₂ = -some (nonsingular_add' h₁ h₂ $ λ _, hy) := -some_add_some_of_Y_ne h₁ h₂ hx hy +some_add_some_of_Y_ne hx hy @[simp] lemma some_add_self_of_Y_ne (hy : y₁ ≠ W.neg_Y x₁ y₁) : some h₁ + some h₁ = some (nonsingular_add h₁ h₁ $ λ _, hy) := -some_add_some_of_Y_ne h₁ h₁ rfl hy +some_add_some_of_Y_ne rfl hy lemma some_add_self_of_Y_ne' (hy : y₁ ≠ W.neg_Y x₁ y₁) : some h₁ + some h₁ = -some (nonsingular_add' h₁ h₁ $ λ _, hy) := -some_add_some_of_Y_ne h₁ h₁ rfl hy +some_add_some_of_Y_ne rfl hy @[simp] lemma some_add_some_of_X_ne (hx : x₁ ≠ x₂) : some h₁ + some h₂ = some (nonsingular_add h₁ h₂ $ λ h, (hx h).elim) := @@ -415,7 +494,7 @@ by rw [← add_def, add, dif_neg hx] lemma some_add_some_of_X_ne' (hx : x₁ ≠ x₂) : some h₁ + some h₂ = -some (nonsingular_add' h₁ h₂ $ λ h, (hx h).elim) := -some_add_some_of_X_ne h₁ h₂ hx +some_add_some_of_X_ne hx end point @@ -431,7 +510,7 @@ namespace point @[simp] lemma add_eq_zero (P Q : W.point) : P + Q = 0 ↔ P = -Q := begin - rcases ⟨P, Q⟩ with ⟨_ | @⟨x₁, y₁, h₁⟩, _ | @⟨x₂, y₂, h₂⟩⟩, + rcases ⟨P, Q⟩ with ⟨_ | @⟨x₁, y₁, _⟩, _ | @⟨x₂, y₂, _⟩⟩, any_goals { refl }, { rw [zero_def, zero_add, eq_neg_iff_eq_neg, neg_zero] }, { simp only [neg_some], @@ -440,11 +519,11 @@ begin by_cases hx : x₁ = x₂, { by_cases hy : y₁ = W.neg_Y x₂ y₂, { exact ⟨hx, hy⟩ }, - { rw [some_add_some_of_Y_ne h₁ h₂ hx hy] at h, + { rw [some_add_some_of_Y_ne hx hy] at h, contradiction } }, - { rw [some_add_some_of_X_ne h₁ h₂ hx] at h, + { rw [some_add_some_of_X_ne hx] at h, contradiction } }, - { exact λ ⟨hx, hy⟩, some_add_some_of_Y_eq h₁ h₂ hx hy } } + { exact λ ⟨hx, hy⟩, some_add_some_of_Y_eq hx hy } } end @[simp] lemma add_left_neg (P : W.point) : -P + P = 0 := by rw [add_eq_zero] @@ -455,6 +534,66 @@ end point end group +section base_change + +/-! ### Nonsingular rational points on a base changed Weierstrass curve -/ + +variables {R : Type u} [comm_ring R] (W : weierstrass_curve R) (F : Type v) [field F] [algebra R F] + (K : Type w) [field K] [algebra R K] [algebra F K] [is_scalar_tower R F K] + +namespace point + +open_locale weierstrass_curve + +/-- The function from `W⟮F⟯` to `W⟮K⟯` induced by a base change from `F` to `K`. -/ +def of_base_change_fun : W⟮F⟯ → W⟮K⟯ +| 0 := 0 +| (some h) := some $ (nonsingular_iff_base_change_of_base_change W F K _ _).mp h + +/-- The group homomorphism from `W⟮F⟯` to `W⟮K⟯` induced by a base change from `F` to `K`. -/ +@[simps] def of_base_change : W⟮F⟯ →+ W⟮K⟯ := +{ to_fun := of_base_change_fun W F K, + map_zero' := rfl, + map_add' := + begin + rintro (_ | @⟨x₁, y₁, _⟩) (_ | @⟨x₂, y₂, _⟩), + any_goals { refl }, + by_cases hx : x₁ = x₂, + { by_cases hy : y₁ = (W.base_change F).neg_Y x₂ y₂, + { simp only [some_add_some_of_Y_eq hx hy, of_base_change_fun], + rw [some_add_some_of_Y_eq $ congr_arg _ hx], + { rw [hy, base_change_neg_Y_of_base_change] } }, + { simp only [some_add_some_of_Y_ne hx hy, of_base_change_fun], + rw [some_add_some_of_Y_ne $ congr_arg _ hx], + { simp only [base_change_add_X_of_base_change, base_change_add_Y_of_base_change, + base_change_slope_of_base_change], + exact ⟨rfl, rfl⟩ }, + { rw [base_change_neg_Y_of_base_change], + contrapose! hy, + exact no_zero_smul_divisors.algebra_map_injective F K hy } } }, + { simp only [some_add_some_of_X_ne hx, of_base_change_fun], + rw [some_add_some_of_X_ne], + { simp only [base_change_add_X_of_base_change, base_change_add_Y_of_base_change, + base_change_slope_of_base_change], + exact ⟨rfl, rfl⟩ }, + { contrapose! hx, + exact no_zero_smul_divisors.algebra_map_injective F K hx } } + end } + +lemma of_base_change_injective : function.injective $ of_base_change W F K := +begin + rintro (_ | _) (_ | _) h, + { refl }, + any_goals { contradiction }, + simp only, + exact ⟨no_zero_smul_divisors.algebra_map_injective F K (some.inj h).left, + no_zero_smul_divisors.algebra_map_injective F K (some.inj h).right⟩ +end + +end point + +end base_change + end weierstrass_curve namespace elliptic_curve diff --git a/src/algebraic_geometry/elliptic_curve/weierstrass.lean b/src/algebraic_geometry/elliptic_curve/weierstrass.lean index b4e8cc54f81e1..70aeba4ce34d4 100644 --- a/src/algebraic_geometry/elliptic_curve/weierstrass.lean +++ b/src/algebraic_geometry/elliptic_curve/weierstrass.lean @@ -85,7 +85,7 @@ private meta def eval_simp : tactic unit := private meta def C_simp : tactic unit := `[simp only [C_0, C_1, C_neg, C_add, C_sub, C_mul, C_pow]] -universes u v +universes u v w variable {R : Type u} @@ -180,12 +180,13 @@ by { dsimp, ring1 } end variable_change +variables (A : Type v) [comm_ring A] [algebra R A] (B : Type w) [comm_ring B] [algebra R B] + [algebra A B] [is_scalar_tower R A B] + section base_change /-! ### Base changes -/ -variables (A : Type v) [comm_ring A] [algebra R A] - /-- The Weierstrass curve over `R` base changed to `A`. -/ @[simps] def base_change : weierstrass_curve A := ⟨algebra_map R A W.a₁, algebra_map R A W.a₂, algebra_map R A W.a₃, algebra_map R A W.a₄, @@ -213,6 +214,11 @@ by { simp only [c₆, base_change_b₂, base_change_b₄, base_change_b₆], map @[simp, nolint simp_nf] lemma base_change_Δ : (W.base_change A).Δ = algebra_map R A W.Δ := by { simp only [Δ, base_change_b₂, base_change_b₄, base_change_b₆, base_change_b₈], map_simp } +lemma base_change_self : W.base_change R = W := by ext; refl + +lemma base_change_base_change : (W.base_change A).base_change B = W.base_change B := +by ext; exact (is_scalar_tower.algebra_map_apply R A B _).symm + end base_change section torsion_polynomial @@ -321,6 +327,20 @@ begin ring1 end +lemma equation_iff_base_change [nontrivial A] [no_zero_smul_divisors R A] (x y : R) : + W.equation x y ↔ (W.base_change A).equation (algebra_map R A x) (algebra_map R A y) := +begin + simp only [equation_iff], + refine ⟨λ h, _, λ h, _⟩, + { convert congr_arg (algebra_map R A) h; { map_simp, refl } }, + { apply no_zero_smul_divisors.algebra_map_injective R A, map_simp, exact h } +end + +lemma equation_iff_base_change_of_base_change [nontrivial B] [no_zero_smul_divisors A B] (x y : A) : + (W.base_change A).equation x y + ↔ (W.base_change B).equation (algebra_map A B x) (algebra_map A B y) := +by rw [equation_iff_base_change (W.base_change A) B, base_change_base_change] + /-! ### Nonsingularity of Weierstrass curves -/ /-- The partial derivative $W_X(X, Y)$ of $W(X, Y)$ with respect to $X$. @@ -376,6 +396,21 @@ begin congr' 4; ring1 end +lemma nonsingular_iff_base_change [nontrivial A] [no_zero_smul_divisors R A] (x y : R) : + W.nonsingular x y ↔ (W.base_change A).nonsingular (algebra_map R A x) (algebra_map R A y) := +begin + rw [nonsingular_iff, nonsingular_iff, and_congr $ W.equation_iff_base_change A x y], + refine ⟨or.imp (not_imp_not.mpr $ λ h, _) (not_imp_not.mpr $ λ h, _), + or.imp (not_imp_not.mpr $ λ h, _) (not_imp_not.mpr $ λ h, _)⟩, + any_goals { apply no_zero_smul_divisors.algebra_map_injective R A, map_simp, exact h }, + any_goals { convert congr_arg (algebra_map R A) h; { map_simp, refl } } +end + +lemma nonsingular_iff_base_change_of_base_change [nontrivial B] [no_zero_smul_divisors A B] + (x y : A) : (W.base_change A).nonsingular x y + ↔ (W.base_change B).nonsingular (algebra_map A B x) (algebra_map A B y) := +by rw [nonsingular_iff_base_change (W.base_change A) B, base_change_base_change] + lemma nonsingular_zero_of_Δ_ne_zero (h : W.equation 0 0) (hΔ : W.Δ ≠ 0) : W.nonsingular 0 0 := by { simp only [equation_zero, nonsingular_zero] at *, contrapose! hΔ, simp [h, hΔ] } From 57ac39bd365c2f80589a700f9fbb664d3a1a30c2 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Mon, 27 Feb 2023 08:41:40 +0000 Subject: [PATCH 0544/1291] chore(data/real/ennreal): drop some lemmas (#18501) Drop lemmas that are in fact more general lemmas applied to `ennreal`. For now, these lemmas are marked as `@[deprecated]` in leanprover-community/mathlib4#2388. --- .../100-theorems-list/30_ballot_problem.lean | 4 +- src/analysis/analytic/composition.lean | 4 +- src/analysis/normed_space/enorm.lean | 2 +- src/analysis/normed_space/lp_space.lean | 10 +-- src/analysis/special_functions/pow.lean | 6 +- src/analysis/specific_limits/basic.lean | 2 +- src/data/real/ennreal.lean | 72 +++++-------------- src/measure_theory/covering/besicovitch.lean | 6 +- .../covering/differentiation.lean | 14 ++-- src/measure_theory/covering/vitali.lean | 2 +- .../conditional_expectation/basic.lean | 18 +++-- .../function/continuous_map_dense.lean | 2 +- .../function/convergence_in_measure.lean | 2 +- src/measure_theory/function/jacobian.lean | 12 ++-- src/measure_theory/function/l2_space.lean | 10 +-- src/measure_theory/function/lp_space.lean | 15 ++-- .../function/simple_func_dense_lp.lean | 6 +- .../function/uniform_integrable.lean | 12 ++-- src/measure_theory/integral/lebesgue.lean | 10 +-- .../integral/mean_inequalities.lean | 11 ++- .../integral/vitali_caratheodory.lean | 7 +- src/measure_theory/measure/doubling.lean | 4 +- src/measure_theory/measure/haar.lean | 2 +- src/measure_theory/measure/haar_lebesgue.lean | 9 ++- src/measure_theory/measure/hausdorff.lean | 21 +++--- src/measure_theory/measure/lebesgue.lean | 4 +- src/measure_theory/measure/measure_space.lean | 3 +- src/measure_theory/measure/regular.lean | 7 +- src/order/filter/ennreal.lean | 3 +- src/probability/integration.lean | 4 +- src/probability/kernel/basic.lean | 2 +- src/probability/martingale/upcrossing.lean | 6 +- src/probability/variance.lean | 14 ++-- src/topology/instances/ennreal.lean | 6 +- .../metric_space/emetric_paracompact.lean | 8 +-- src/topology/metric_space/emetric_space.lean | 4 +- .../metric_space/hausdorff_dimension.lean | 2 +- .../metric_space/hausdorff_distance.lean | 6 +- .../metric_space/metric_separated.lean | 2 +- 39 files changed, 145 insertions(+), 189 deletions(-) diff --git a/archive/100-theorems-list/30_ballot_problem.lean b/archive/100-theorems-list/30_ballot_problem.lean index 1ba4b7e1f4299..a72c2e62f322e 100644 --- a/archive/100-theorems-list/30_ballot_problem.lean +++ b/archive/100-theorems-list/30_ballot_problem.lean @@ -425,9 +425,9 @@ begin { rw ballot_problem' q p qp, rw [ennreal.to_real_div, ← nat.cast_add, ← nat.cast_add, ennreal.to_real_nat, ennreal.to_real_sub_of_le, ennreal.to_real_nat, ennreal.to_real_nat], - exacts [ennreal.coe_nat_le_coe_nat.2 qp.le, ennreal.nat_ne_top _] }, + exacts [nat.cast_le.2 qp.le, ennreal.nat_ne_top _] }, rwa ennreal.to_real_eq_to_real (measure_lt_top _ _).ne at this, - { simp only [ne.def, ennreal.div_eq_top, tsub_eq_zero_iff_le, ennreal.coe_nat_le_coe_nat, + { simp only [ne.def, ennreal.div_eq_top, tsub_eq_zero_iff_le, nat.cast_le, not_le, add_eq_zero_iff, nat.cast_eq_zero, ennreal.add_eq_top, ennreal.nat_ne_top, or_self, not_false_iff, and_true], push_neg, diff --git a/src/analysis/analytic/composition.lean b/src/analysis/analytic/composition.lean index a0e417d0692ad..37f8e87141756 100644 --- a/src/analysis/analytic/composition.lean +++ b/src/analysis/analytic/composition.lean @@ -444,8 +444,8 @@ 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 fact that there are `2^(n-1)` compositions of `n`, giving at most a geometric loss. -/ - rcases ennreal.lt_iff_exists_nnreal_btwn.1 (lt_min 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⟩, + rcases ennreal.lt_iff_exists_nnreal_btwn.1 (lt_min zero_lt_one hq) with ⟨rq, rq_pos, hrq⟩, + rcases ennreal.lt_iff_exists_nnreal_btwn.1 (lt_min zero_lt_one hp) with ⟨rp, rp_pos, hrp⟩, 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, ‖q n‖₊ * rq^n ≤ Cq := q.nnnorm_mul_pow_le_of_lt_radius hrq.2, diff --git a/src/analysis/normed_space/enorm.lean b/src/analysis/normed_space/enorm.lean index 5326b4c7d8729..b355dfce88d26 100644 --- a/src/analysis/normed_space/enorm.lean +++ b/src/analysis/normed_space/enorm.lean @@ -68,7 +68,7 @@ begin 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' _ _) }, + { exact mul_le_mul_left' (e.map_smul_le' _ _) _ }, { rw [← mul_assoc, nnnorm_inv, ennreal.coe_inv, ennreal.mul_inv_cancel _ ennreal.coe_ne_top, one_mul]; simp [hc] } end diff --git a/src/analysis/normed_space/lp_space.lean b/src/analysis/normed_space/lp_space.lean index 158f95d00a53c..77c3b45e3042f 100644 --- a/src/analysis/normed_space/lp_space.lean +++ b/src/analysis/normed_space/lp_space.lean @@ -610,7 +610,7 @@ end instance [fact (1 ≤ p)] : normed_space 𝕜 (lp E p) := { norm_smul_le := λ c f, begin - have hp : 0 < p := ennreal.zero_lt_one.trans_le (fact.out _), + have hp : 0 < p := zero_lt_one.trans_le (fact.out _), simp [norm_const_smul hp.ne'] end } @@ -944,7 +944,7 @@ by linarith [lp.norm_sub_norm_compl_sub_single hp f s] protected lemma has_sum_single [fact (1 ≤ p)] (hp : p ≠ ⊤) (f : lp E p) : has_sum (λ i : α, lp.single p i (f i : E i)) f := begin - have hp₀ : 0 < p := ennreal.zero_lt_one.trans_le (fact.out _), + have hp₀ : 0 < p := zero_lt_one.trans_le (fact.out _), have hp' : 0 < p.to_real := ennreal.to_real_pos hp₀.ne' hp, have := lp.has_sum_norm hp' f, rw [has_sum, metric.tendsto_nhds] at this ⊢, @@ -975,7 +975,7 @@ open_locale topology uniformity /-- The coercion from `lp E p` to `Π i, E i` is uniformly continuous. -/ lemma uniform_continuous_coe [_i : fact (1 ≤ p)] : uniform_continuous (coe : lp E p → Π i, E i) := begin - have hp : p ≠ 0 := (ennreal.zero_lt_one.trans_le _i.elim).ne', + have hp : p ≠ 0 := (zero_lt_one.trans_le _i.elim).ne', rw uniform_continuous_pi, intros i, rw normed_add_comm_group.uniformity_basis_dist.uniform_continuous_iff @@ -1008,7 +1008,7 @@ lemma sum_rpow_le_of_tendsto (hp : p ≠ ∞) {C : ℝ} {F : ι → lp E p} (hCF {f : Π a, E a} (hf : tendsto (id (λ i, F i) : ι → Π a, E a) l (𝓝 f)) (s : finset α) : ∑ (i : α) in s, ‖f i‖ ^ p.to_real ≤ C ^ p.to_real := begin - have hp' : p ≠ 0 := (ennreal.zero_lt_one.trans_le _i.elim).ne', + have hp' : p ≠ 0 := (zero_lt_one.trans_le _i.elim).ne', have hp'' : 0 < p.to_real := ennreal.to_real_pos hp' hp, let G : (Π a, E a) → ℝ := λ f, ∑ a in s, ‖f a‖ ^ p.to_real, have hG : continuous G, @@ -1034,7 +1034,7 @@ begin unfreezingI { rcases eq_top_or_lt_top p with rfl | hp }, { apply norm_le_of_forall_le hC, exact norm_apply_le_of_tendsto hCF hf, }, - { have : 0 < p := ennreal.zero_lt_one.trans_le _i.elim, + { have : 0 < p := zero_lt_one.trans_le _i.elim, have hp' : 0 < p.to_real := ennreal.to_real_pos this.ne' hp.ne, apply norm_le_of_forall_sum_le hp' hC, exact sum_rpow_le_of_tendsto hp.ne hCF hf, } diff --git a/src/analysis/special_functions/pow.lean b/src/analysis/special_functions/pow.lean index 3aec5d760b620..a91dee0dfda9f 100644 --- a/src/analysis/special_functions/pow.lean +++ b/src/analysis/special_functions/pow.lean @@ -1941,7 +1941,7 @@ end lemma rpow_pos_of_nonneg {p : ℝ} {x : ℝ≥0∞} (hx_pos : 0 < x) (hp_nonneg : 0 ≤ p) : 0 < x^p := begin by_cases hp_zero : p = 0, - { simp [hp_zero, ennreal.zero_lt_one], }, + { simp [hp_zero, zero_lt_one], }, { rw ←ne.def at hp_zero, have hp_pos := lt_of_le_of_ne hp_nonneg hp_zero.symm, rw ←zero_rpow_of_pos hp_pos, exact rpow_lt_rpow hx_pos hp_pos, }, @@ -1972,7 +1972,7 @@ end lemma rpow_lt_one_of_one_lt_of_neg {x : ℝ≥0∞} {z : ℝ} (hx : 1 < x) (hz : z < 0) : x^z < 1 := begin cases x, - { simp [top_rpow_of_neg hz, ennreal.zero_lt_one] }, + { simp [top_rpow_of_neg hz, zero_lt_one] }, { simp only [some_eq_coe, one_lt_coe_iff] at hx, simp [coe_rpow_of_ne_zero (ne_of_gt (lt_trans zero_lt_one hx)), nnreal.rpow_lt_one_of_one_lt_of_neg hx hz] }, @@ -1981,7 +1981,7 @@ end lemma rpow_le_one_of_one_le_of_neg {x : ℝ≥0∞} {z : ℝ} (hx : 1 ≤ x) (hz : z < 0) : x^z ≤ 1 := begin cases x, - { simp [top_rpow_of_neg hz, ennreal.zero_lt_one] }, + { simp [top_rpow_of_neg hz, zero_lt_one] }, { simp only [one_le_coe_iff, some_eq_coe] at hx, simp [coe_rpow_of_ne_zero (ne_of_gt (lt_of_lt_of_le zero_lt_one hx)), nnreal.rpow_le_one_of_one_le_of_nonpos hx (le_of_lt hz)] }, diff --git a/src/analysis/specific_limits/basic.lean b/src/analysis/specific_limits/basic.lean index 28f25521d4b77..2230dbead1b92 100644 --- a/src/analysis/specific_limits/basic.lean +++ b/src/analysis/specific_limits/basic.lean @@ -506,7 +506,7 @@ begin refine ⟨λ i, δ' i / max 1 (w i), λ i, div_pos (Hpos _) (this i), _⟩, refine lt_of_le_of_lt (ennreal.tsum_le_tsum $ λ i, _) Hsum, rw [coe_div (this i).ne'], - refine mul_le_of_le_div' (ennreal.mul_le_mul le_rfl $ ennreal.inv_le_inv.2 _), + refine mul_le_of_le_div' (mul_le_mul_left' (ennreal.inv_le_inv.2 _) _), exact coe_le_coe.2 (le_max_right _ _) end diff --git a/src/data/real/ennreal.lean b/src/data/real/ennreal.lean index c35139f2da10b..e0597878dbada 100644 --- a/src/data/real/ennreal.lean +++ b/src/data/real/ennreal.lean @@ -80,7 +80,7 @@ variables {α : Type*} {β : Type*} semilattice_sup, distrib_lattice, order_bot, bounded_order, canonically_ordered_comm_semiring, complete_linear_order, densely_ordered, nontrivial, canonically_linear_ordered_add_monoid, has_sub, has_ordered_sub, - linear_ordered_add_comm_monoid_with_top]] + linear_ordered_add_comm_monoid_with_top, char_zero]] def ennreal := with_top ℝ≥0 localized "notation (name := ennreal) `ℝ≥0∞` := ennreal" in ennreal @@ -222,7 +222,6 @@ lemma coe_mono : monotone (coe : ℝ≥0 → ℝ≥0∞) := λ _ _, coe_le_coe.2 @[simp, norm_cast] lemma zero_eq_coe : 0 = (↑r : ℝ≥0∞) ↔ 0 = r := coe_eq_coe @[simp, norm_cast] lemma coe_eq_one : (↑r : ℝ≥0∞) = 1 ↔ r = 1 := coe_eq_coe @[simp, norm_cast] lemma one_eq_coe : 1 = (↑r : ℝ≥0∞) ↔ 1 = r := coe_eq_coe -@[simp] lemma coe_nonneg : 0 ≤ (↑r : ℝ≥0∞) := coe_le_coe.2 $ zero_le _ @[simp, norm_cast] lemma coe_pos : 0 < (↑r : ℝ≥0∞) ↔ 0 < r := coe_lt_coe lemma coe_ne_zero : (r : ℝ≥0∞) ≠ 0 ↔ r ≠ 0 := not_congr coe_eq_coe @@ -255,12 +254,9 @@ lemma to_real_eq_to_real_iff' {x y : ℝ≥0∞} (hx : x ≠ ⊤) (hy : y ≠ x.to_real = y.to_real ↔ x = y := by simp only [ennreal.to_real, nnreal.coe_eq, to_nnreal_eq_to_nnreal_iff' hx hy] -protected lemma zero_lt_one : 0 < (1 : ℝ≥0∞) := zero_lt_one - @[simp] lemma one_lt_two : (1 : ℝ≥0∞) < 2 := coe_one ▸ coe_two ▸ by exact_mod_cast (one_lt_two : 1 < 2) -@[simp] lemma zero_lt_two : (0 : ℝ≥0∞) < 2 := lt_trans zero_lt_one one_lt_two -lemma two_ne_zero : (2 : ℝ≥0∞) ≠ 0 := (ne_of_lt zero_lt_two).symm + lemma two_ne_top : (2 : ℝ≥0∞) ≠ ∞ := coe_two ▸ coe_ne_top /-- `(1 : ℝ≥0∞) ≤ 1`, recorded as a `fact` for use with `Lp` spaces. -/ @@ -301,9 +297,6 @@ lemma supr_ennreal {α : Type*} [complete_lattice α] {f : ℝ≥0∞ → α} : (⨆ n, f n) = (⨆ n : ℝ≥0, f n) ⊔ f ∞ := @infi_ennreal αᵒᵈ _ _ -@[simp] lemma add_top : a + ∞ = ∞ := add_top _ -@[simp] lemma top_add : ∞ + a = ∞ := top_add _ - /-- Coercion `ℝ≥0 → ℝ≥0∞` as a `ring_hom`. -/ def of_nnreal_hom : ℝ≥0 →+* ℝ≥0∞ := ⟨coe, coe_one, λ _ _, coe_mul, coe_zero, λ _ _, coe_add⟩ @@ -486,21 +479,6 @@ end ↑(s.sup f) = s.sup (λ x, (f x : ℝ≥0∞)) := finset.comp_sup_eq_sup_comp_of_is_total _ coe_mono rfl -lemma pow_le_pow {n m : ℕ} (ha : 1 ≤ a) (h : n ≤ m) : a ^ n ≤ a ^ m := -begin - cases a, - { cases m, - { rw eq_bot_iff.mpr h, - exact le_rfl }, - { rw [none_eq_top, top_pow (nat.succ_pos m)], - exact le_top } }, - { rw [some_eq_coe, ← coe_pow, ← coe_pow, coe_le_coe], - exact pow_le_pow (by simpa using ha) h } -end - -lemma one_le_pow_of_one_le (ha : 1 ≤ a) (n : ℕ) : 1 ≤ a ^ n := -by simpa using pow_le_pow ha (zero_le n) - @[simp] lemma max_eq_zero_iff : max a b = 0 ↔ a = 0 ∧ b = 0 := by simp only [nonpos_iff_eq_zero.symm, max_le_iff] @@ -592,15 +570,11 @@ begin exact tsub_add_cancel_of_le ad.le end -lemma coe_nat_lt_coe {n : ℕ} : (n : ℝ≥0∞) < r ↔ ↑n < r := ennreal.coe_nat n ▸ coe_lt_coe -lemma coe_lt_coe_nat {n : ℕ} : (r : ℝ≥0∞) < n ↔ r < n := ennreal.coe_nat n ▸ coe_lt_coe -@[simp, norm_cast] lemma coe_nat_lt_coe_nat {m n : ℕ} : (m : ℝ≥0∞) < n ↔ m < n := -ennreal.coe_nat n ▸ coe_nat_lt_coe.trans nat.cast_lt -lemma coe_nat_mono : strict_mono (coe : ℕ → ℝ≥0∞) := λ _ _, coe_nat_lt_coe_nat.2 -@[simp, norm_cast] lemma coe_nat_le_coe_nat {m n : ℕ} : (m : ℝ≥0∞) ≤ n ↔ m ≤ n := -coe_nat_mono.le_iff_le +@[simp, norm_cast] lemma coe_nat_lt_coe {n : ℕ} : (n : ℝ≥0∞) < r ↔ ↑n < r := +ennreal.coe_nat n ▸ coe_lt_coe -instance : char_zero ℝ≥0∞ := ⟨coe_nat_mono.injective⟩ +@[simp, norm_cast] lemma coe_lt_coe_nat {n : ℕ} : (r : ℝ≥0∞) < n ↔ r < n := +ennreal.coe_nat n ▸ coe_lt_coe protected lemma exists_nat_gt {r : ℝ≥0∞} (h : r ≠ ∞) : ∃n:ℕ, r < n := begin @@ -674,9 +648,6 @@ section complete_lattice lemma coe_Sup {s : set ℝ≥0} : bdd_above s → (↑(Sup s) : ℝ≥0∞) = (⨆a∈s, ↑a) := with_top.coe_Sup lemma coe_Inf {s : set ℝ≥0} : s.nonempty → (↑(Inf s) : ℝ≥0∞) = (⨅a∈s, ↑a) := with_top.coe_Inf -@[simp] lemma top_mem_upper_bounds {s : set ℝ≥0∞} : ∞ ∈ upper_bounds s := -assume x hx, le_top - lemma coe_mem_upper_bounds {s : set ℝ≥0} : ↑r ∈ upper_bounds ((coe : ℝ≥0 → ℝ≥0∞) '' s) ↔ r ∈ upper_bounds s := by simp [upper_bounds, ball_image_iff, -mem_image, *] {contextual := tt} @@ -685,9 +656,6 @@ end complete_lattice section mul -@[mono] lemma mul_le_mul : a ≤ b → c ≤ d → a * c ≤ b * d := -mul_le_mul' - @[mono] lemma mul_lt_mul (ac : a < c) (bd : b < d) : a * b < c * d := begin rcases lt_iff_exists_nnreal_btwn.1 ac with ⟨a', aa', a'c⟩, @@ -698,12 +666,14 @@ begin calc ↑(a * b) < ↑(a' * b') : coe_lt_coe.2 (mul_lt_mul' aa'.le bb' (zero_le _) ((zero_le a).trans_lt aa')) ... = ↑a' * ↑b' : coe_mul - ... ≤ c * d : mul_le_mul a'c.le b'd.le + ... ≤ c * d : mul_le_mul' a'c.le b'd.le end -lemma mul_left_mono : monotone ((*) a) := λ b c, mul_le_mul le_rfl +-- TODO: generalize to `covariant_class α α (*) (≤)` +lemma mul_left_mono : monotone ((*) a) := λ b c, mul_le_mul' le_rfl -lemma mul_right_mono : monotone (λ x, x * a) := λ b c h, mul_le_mul h le_rfl +-- TODO: generalize to `covariant_class α α (swap (*)) (≤)` +lemma mul_right_mono : monotone (λ x, x * a) := λ b c h, mul_le_mul' h le_rfl lemma pow_strict_mono {n : ℕ} (hn : n ≠ 0) : strict_mono (λ (x : ℝ≥0∞), x^n) := begin @@ -727,7 +697,7 @@ begin intros x y h, contrapose! h, simpa only [← mul_assoc, ← coe_mul, inv_mul_cancel h0, coe_one, one_mul] - using mul_le_mul' (le_refl ↑a⁻¹) h + using mul_le_mul_left' h (↑a⁻¹) end lemma mul_eq_mul_left (h₀ : a ≠ 0) (hinf : a ≠ ∞) : a * b = a * c ↔ b = c := @@ -974,7 +944,7 @@ lemma bit0_injective : function.injective (bit0 : ℝ≥0∞ → ℝ≥0∞) := @[simp] lemma bit0_inj : bit0 a = bit0 b ↔ a = b := bit0_injective.eq_iff @[simp] lemma bit0_eq_zero_iff : bit0 a = 0 ↔ a = 0 := bit0_injective.eq_iff' bit0_zero -@[simp] lemma bit0_top : bit0 ∞ = ∞ := add_top +@[simp] lemma bit0_top : bit0 ∞ = ∞ := add_top _ @[simp] lemma bit0_eq_top_iff : bit0 a = ∞ ↔ a = ∞ := bit0_injective.eq_iff' bit0_top @[mono] lemma bit1_strict_mono : strict_mono (bit1 : ℝ≥0∞ → ℝ≥0∞) := @@ -1151,12 +1121,6 @@ def _root_.order_iso.inv_ennreal : ℝ≥0∞ ≃o ℝ≥0∞ᵒᵈ := lemma _root_.order_iso.inv_ennreal_symm_apply : order_iso.inv_ennreal.symm a = (order_dual.of_dual a)⁻¹ := rfl -protected lemma pow_le_pow_of_le_one {n m : ℕ} (ha : a ≤ 1) (h : n ≤ m) : a ^ m ≤ a ^ n := -begin - rw [←inv_inv a, ← ennreal.inv_pow, ← @ennreal.inv_pow a⁻¹, ennreal.inv_le_inv], - exact pow_le_pow (ennreal.one_le_inv.2 ha) h -end - @[simp] lemma div_top : a / ∞ = 0 := by rw [div_eq_mul_inv, inv_top, mul_zero] @[simp] lemma top_div_coe : ∞ / p = ∞ := by simp [div_eq_mul_inv, top_mul] @@ -1176,7 +1140,7 @@ lemma div_eq_top : a / b = ∞ ↔ (a ≠ 0 ∧ b = 0) ∨ (a = ∞ ∧ b ≠ by simp [div_eq_mul_inv, ennreal.mul_eq_top] protected lemma le_div_iff_mul_le (h0 : b ≠ 0 ∨ c ≠ 0) (ht : b ≠ ∞ ∨ c ≠ ∞) : - a ≤ c / b ↔ a * b ≤ c := + a ≤ c / b ↔ a * b ≤ c := begin induction b using with_top.rec_top_coe, { lift c to ℝ≥0 using ht.neg_resolve_left rfl, @@ -1239,7 +1203,7 @@ end by rw [← one_div, ennreal.le_div_iff_mul_le]; { right, simp } protected lemma div_le_div (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) +div_eq_mul_inv b d ▸ div_eq_mul_inv a c ▸ mul_le_mul' hab (ennreal.inv_le_inv.mpr hdc) protected lemma div_le_div_left (h : a ≤ b) (c : ℝ≥0∞) : c / b ≤ c / a := ennreal.div_le_div le_rfl h @@ -1389,8 +1353,8 @@ lemma exists_nat_pos_mul_gt (ha : a ≠ 0) (hb : b ≠ ∞) : begin have : b / a ≠ ∞, from mul_ne_top hb (inv_ne_top.2 ha), refine (ennreal.exists_nat_gt this).imp (λ n hn, _), - have : ↑0 < (n : ℝ≥0∞), from lt_of_le_of_lt (by simp) hn, - refine ⟨coe_nat_lt_coe_nat.1 this, _⟩, + have : 0 < (n : ℝ≥0∞), from lt_of_le_of_lt (zero_le _) hn, + refine ⟨nat.cast_pos.1 this, _⟩, rwa [← ennreal.div_lt_iff (or.inl ha) (or.inr hb)] end @@ -1504,7 +1468,7 @@ begin exact lt_of_lt_of_le (int.neg_succ_lt_zero _) (int.of_nat_nonneg _) }, { simp only [zpow_neg_succ_of_nat, int.of_nat_eq_coe, zpow_coe_nat], refine (ennreal.inv_le_one.2 _).trans _; - exact ennreal.one_le_pow_of_one_le hx _, }, + exact one_le_pow_of_one_le' hx _, }, { simp only [zpow_neg_succ_of_nat, ennreal.inv_le_inv], apply pow_le_pow hx, simpa only [←int.coe_nat_le_coe_nat_iff, neg_le_neg_iff, int.coe_nat_add, int.coe_nat_one, diff --git a/src/measure_theory/covering/besicovitch.lean b/src/measure_theory/covering/besicovitch.lean index 86423e2f90bc3..775b5ca0256da 100644 --- a/src/measure_theory/covering/besicovitch.lean +++ b/src/measure_theory/covering/besicovitch.lean @@ -565,7 +565,7 @@ begin apply (ennreal.mul_lt_mul_left hμs.ne' (measure_lt_top μ s).ne).2, rw ennreal.inv_lt_inv, conv_lhs {rw ← add_zero (N : ℝ≥0∞) }, - exact ennreal.add_lt_add_left (ennreal.nat_ne_top N) ennreal.zero_lt_one }, + exact ennreal.add_lt_add_left (ennreal.nat_ne_top N) zero_lt_one }, have B : μ (o ∩ v i) = ∑' (x : u i), μ (o ∩ closed_ball x (r x)), { have : o ∩ v i = ⋃ (x : s) (hx : x ∈ u i), o ∩ closed_ball x (r x), by simp only [inter_Union], rw [this, measure_bUnion (u_count i)], @@ -750,13 +750,13 @@ begin ≤ (N/(N+1)) * μ (s \ ⋃ (p : α × ℝ) (hp : p ∈ u n), closed_ball p.fst p.snd) : by { rw u_succ, exact (hF (u n) (Pu n)).2.2 } ... ≤ (N/(N+1))^n.succ * μ s : - by { rw [pow_succ, mul_assoc], exact ennreal.mul_le_mul le_rfl IH } }, + by { rw [pow_succ, mul_assoc], exact mul_le_mul_left' IH _ } }, have C : tendsto (λ (n : ℕ), ((N : ℝ≥0∞)/(N+1))^n * μ s) at_top (𝓝 (0 * μ s)), { apply ennreal.tendsto.mul_const _ (or.inr (measure_lt_top μ s).ne), apply ennreal.tendsto_pow_at_top_nhds_0_of_lt_1, rw [ennreal.div_lt_iff, one_mul], { conv_lhs {rw ← add_zero (N : ℝ≥0∞) }, - exact ennreal.add_lt_add_left (ennreal.nat_ne_top N) ennreal.zero_lt_one }, + exact ennreal.add_lt_add_left (ennreal.nat_ne_top N) zero_lt_one }, { simp only [true_or, add_eq_zero_iff, ne.def, not_false_iff, one_ne_zero, and_false] }, { simp only [ennreal.nat_ne_top, ne.def, not_false_iff, or_true] } }, rw zero_mul at C, diff --git a/src/measure_theory/covering/differentiation.lean b/src/measure_theory/covering/differentiation.lean index e70665357ccb9..8dd8cd38b8c53 100644 --- a/src/measure_theory/covering/differentiation.lean +++ b/src/measure_theory/covering/differentiation.lean @@ -178,14 +178,14 @@ begin rw [ennreal.mul_inv_cancel (ennreal.coe_pos.2 εpos).ne' ennreal.coe_ne_top, one_mul], end ... ≤ ε⁻¹ * ρ (s ∩ o) : begin - apply ennreal.mul_le_mul le_rfl, + refine mul_le_mul_left' _ _, refine v.measure_le_of_frequently_le ρ ((measure.absolutely_continuous.refl μ).smul ε) _ _, assume x hx, rw hs at hx, simp only [mem_inter_iff, not_lt, not_eventually, mem_set_of_eq] at hx, exact hx.1 end - ... ≤ ε⁻¹ * ρ o : ennreal.mul_le_mul le_rfl (measure_mono (inter_subset_right _ _)) + ... ≤ ε⁻¹ * ρ o : mul_le_mul_left' (measure_mono (inter_subset_right _ _)) _ ... = 0 : by rw [ρo, mul_zero] }, obtain ⟨u, u_anti, u_pos, u_lim⟩ : ∃ (u : ℕ → ℝ≥0), strict_anti u ∧ (∀ (n : ℕ), 0 < u n) ∧ tendsto u at_top (𝓝 0) := @@ -202,7 +202,7 @@ begin filter_upwards [hx n, h'x, v.eventually_measure_lt_top x], assume a ha μa_pos μa_lt_top, rw ennreal.div_lt_iff (or.inl μa_pos.ne') (or.inl μa_lt_top.ne), - exact ha.trans_le (ennreal.mul_le_mul ((ennreal.coe_le_coe.2 hn.le).trans w_lt.le) le_rfl) + exact ha.trans_le (mul_le_mul_right' ((ennreal.coe_le_coe.2 hn.le).trans w_lt.le) _) end section absolutely_continuous @@ -453,7 +453,7 @@ begin ... ≤ p * μ (s ∩ t) + 0 : add_le_add H ((measure_mono (inter_subset_right _ _)).trans (hρ A).le) ... ≤ p * μ s : - by { rw add_zero, exact ennreal.mul_le_mul le_rfl (measure_mono (inter_subset_left _ _)) }, + by { rw add_zero, exact mul_le_mul_left' (measure_mono (inter_subset_left _ _)) _ }, refine v.measure_le_of_frequently_le _ hρ _ (λ x hx, _), have I : ∀ᶠ (b : set α) in v.filter_at x, ρ b / μ b < p := (tendsto_order.1 hx.2).2 _ (h hx.1), apply I.frequently.mono (λ a ha, _), @@ -477,7 +477,7 @@ begin ... ≤ ρ (s ∩ t) + q * μ tᶜ : begin apply add_le_add H, rw [coe_nnreal_smul_apply], - exact ennreal.mul_le_mul le_rfl (measure_mono (inter_subset_right _ _)), + exact mul_le_mul_left' (measure_mono (inter_subset_right _ _)) _, end ... ≤ ρ s : by { rw [A, mul_zero, add_zero], exact measure_mono (inter_subset_left _ _) }, @@ -584,7 +584,7 @@ begin abel, end ... ≤ t^2 * ρ (s ∩ f ⁻¹' I) : begin - apply ennreal.mul_le_mul le_rfl _, + refine mul_le_mul_left' _ _, rw ← ennreal.coe_zpow (zero_lt_one.trans ht).ne', apply v.mul_measure_le_of_subset_lt_lim_ratio_meas hρ, assume x hx, @@ -646,7 +646,7 @@ begin ... ≤ ∫⁻ x in s ∩ f⁻¹' I, t * f x ∂μ : begin apply lintegral_mono_ae ((ae_restrict_iff' M).2 (eventually_of_forall (λ x hx, _))), rw [add_comm, ennreal.zpow_add t_ne_zero ennreal.coe_ne_top, zpow_one], - exact ennreal.mul_le_mul le_rfl hx.2.1, + exact mul_le_mul_left' hx.2.1 _, end ... = t * ∫⁻ x in s ∩ f⁻¹' I, f x ∂μ : lintegral_const_mul _ f_meas }, calc ρ s = ρ (s ∩ f⁻¹' {0}) + ρ (s ∩ f⁻¹' {∞}) + ∑' (n : ℤ), ρ (s ∩ f⁻¹' (Ico (t^n) (t^(n+1)))) : diff --git a/src/measure_theory/covering/vitali.lean b/src/measure_theory/covering/vitali.lean index 00142dc43e186..16df444849fe1 100644 --- a/src/measure_theory/covering/vitali.lean +++ b/src/measure_theory/covering/vitali.lean @@ -388,7 +388,7 @@ begin measure_Union_le _ ... ≤ ∑' (a : {a // a ∉ w}), C * μ (B a) : ennreal.tsum_le_tsum (λ a, μB a (ut (vu a.1.2))) ... = C * ∑' (a : {a // a ∉ w}), μ (B a) : ennreal.tsum_mul_left - ... ≤ C * (ε / C) : ennreal.mul_le_mul le_rfl hw.le + ... ≤ C * (ε / C) : mul_le_mul_left' hw.le _ ... ≤ ε : ennreal.mul_div_le end diff --git a/src/measure_theory/function/conditional_expectation/basic.lean b/src/measure_theory/function/conditional_expectation/basic.lean index 4ad8e7062a3e9..46a556bb163ba 100644 --- a/src/measure_theory/function/conditional_expectation/basic.lean +++ b/src/measure_theory/function/conditional_expectation/basic.lean @@ -1081,7 +1081,7 @@ begin { refine mem_ℒp.const_inner _ _, rw Lp_meas_coe, exact Lp.mem_ℒp _, }, have h_eq : h_mem_Lp.to_Lp _ =ᵐ[μ] λ a, ⟪c, condexp_L2 𝕜 hm f a⟫, from h_mem_Lp.coe_fn_to_Lp, refine eventually_eq.trans _ h_eq, - refine Lp.ae_eq_of_forall_set_integral_eq' hm _ _ ennreal.zero_lt_two.ne.symm ennreal.coe_ne_top + refine Lp.ae_eq_of_forall_set_integral_eq' hm _ _ two_ne_zero ennreal.coe_ne_top (λ s hs hμs, integrable_on_condexp_L2_of_measure_ne_top hm hμs.ne _) _ _ _ _, { intros s hs hμs, rw [integrable_on, integrable_congr (ae_restrict_of_ae h_eq)], @@ -1129,7 +1129,7 @@ variables (𝕜 𝕜') lemma condexp_L2_comp_continuous_linear_map (hm : m ≤ m0) (T : E' →L[ℝ] E'') (f : α →₂[μ] E') : (condexp_L2 𝕜' hm (T.comp_Lp f) : α →₂[μ] E'') =ᵐ[μ] T.comp_Lp (condexp_L2 𝕜 hm f : α →₂[μ] E') := begin - refine Lp.ae_eq_of_forall_set_integral_eq' hm _ _ ennreal.zero_lt_two.ne.symm ennreal.coe_ne_top + refine Lp.ae_eq_of_forall_set_integral_eq' hm _ _ two_ne_zero ennreal.coe_ne_top (λ s hs hμs, integrable_on_condexp_L2_of_measure_ne_top hm hμs.ne _) (λ s hs hμs, integrable_on_Lp_of_measure_ne_top _ fact_one_le_two_ennreal.elim hμs.ne) _ _ _, @@ -1197,7 +1197,7 @@ begin exact (Lp.strongly_measurable _).ennnorm end ... ≤ μ (s ∩ t) * ‖x‖₊ : - ennreal.mul_le_mul (lintegral_nnnorm_condexp_L2_indicator_le_real hs hμs ht hμt) le_rfl + mul_le_mul_right' (lintegral_nnnorm_condexp_L2_indicator_le_real hs hμs ht hμt) _ lemma lintegral_nnnorm_condexp_L2_indicator_le (hm : m ≤ m0) (hs : measurable_set s) (hμs : μ s ≠ ∞) (x : E') [sigma_finite (μ.trim hm)] : @@ -1207,8 +1207,7 @@ begin { rw Lp_meas_coe, exact (Lp.ae_strongly_measurable _).ennnorm }, refine (set_lintegral_nnnorm_condexp_L2_indicator_le hm hs hμs x ht hμt).trans _, - refine ennreal.mul_le_mul _ le_rfl, - exact measure_mono (set.inter_subset_left _ _), + exact mul_le_mul_right' (measure_mono (set.inter_subset_left _ _)) _ end /-- If the measure `μ.trim hm` is sigma-finite, then the conditional expectation of a measurable set @@ -1221,7 +1220,7 @@ begin (ennreal.mul_lt_top hμs ennreal.coe_ne_top) _ _, { rw Lp_meas_coe, exact Lp.ae_strongly_measurable _, }, { refine λ t ht hμt, (set_lintegral_nnnorm_condexp_L2_indicator_le hm hs hμs x ht hμt).trans _, - exact ennreal.mul_le_mul (measure_mono (set.inter_subset_left _ _)) le_rfl, }, + exact mul_le_mul_right' (measure_mono (set.inter_subset_left _ _)) _, }, end end condexp_L2_indicator @@ -1284,7 +1283,7 @@ begin exact (Lp.strongly_measurable _).ennnorm end ... ≤ μ (s ∩ t) * ‖x‖₊ : - ennreal.mul_le_mul (lintegral_nnnorm_condexp_L2_indicator_le_real hs hμs ht hμt) le_rfl + mul_le_mul_right' (lintegral_nnnorm_condexp_L2_indicator_le_real hs hμs ht hμt) _ lemma lintegral_nnnorm_condexp_ind_smul_le (hm : m ≤ m0) (hs : measurable_set s) (hμs : μ s ≠ ∞) (x : G) [sigma_finite (μ.trim hm)] : @@ -1293,8 +1292,7 @@ begin refine lintegral_le_of_forall_fin_meas_le' hm (μ s * ‖x‖₊) _ (λ t ht hμt, _), { exact (Lp.ae_strongly_measurable _).ennnorm }, refine (set_lintegral_nnnorm_condexp_ind_smul_le hm hs hμs x ht hμt).trans _, - refine ennreal.mul_le_mul _ le_rfl, - exact measure_mono (set.inter_subset_left _ _), + exact mul_le_mul_right' (measure_mono (set.inter_subset_left _ _)) _ end /-- If the measure `μ.trim hm` is sigma-finite, then the conditional expectation of a measurable set @@ -1307,7 +1305,7 @@ begin (ennreal.mul_lt_top hμs ennreal.coe_ne_top) _ _, { exact Lp.ae_strongly_measurable _, }, { refine λ t ht hμt, (set_lintegral_nnnorm_condexp_ind_smul_le hm hs hμs x ht hμt).trans _, - exact ennreal.mul_le_mul (measure_mono (set.inter_subset_left _ _)) le_rfl, }, + exact mul_le_mul_right' (measure_mono (set.inter_subset_left _ _)) _, }, end lemma condexp_ind_smul_empty {x : G} : diff --git a/src/measure_theory/function/continuous_map_dense.lean b/src/measure_theory/function/continuous_map_dense.lean index b01c24e6cec64..de100258d3fd3 100644 --- a/src/measure_theory/function/continuous_map_dense.lean +++ b/src/measure_theory/function/continuous_map_dense.lean @@ -59,7 +59,7 @@ variables [normed_space ℝ E] lemma bounded_continuous_function_dense [μ.weakly_regular] : (bounded_continuous_function E p μ).topological_closure = ⊤ := begin - have hp₀ : 0 < p := lt_of_lt_of_le ennreal.zero_lt_one _i.elim, + have hp₀ : 0 < p := lt_of_lt_of_le zero_lt_one _i.elim, have hp₀' : 0 ≤ 1 / p.to_real := div_nonneg zero_le_one ennreal.to_real_nonneg, have hp₀'' : 0 < p.to_real, { simpa [← ennreal.to_real_lt_to_real ennreal.zero_ne_top hp] using hp₀ }, diff --git a/src/measure_theory/function/convergence_in_measure.lean b/src/measure_theory/function/convergence_in_measure.lean index 562f06d6eeb93..01d66d182cc2b 100644 --- a/src/measure_theory/function/convergence_in_measure.lean +++ b/src/measure_theory/function/convergence_in_measure.lean @@ -367,7 +367,7 @@ end lemma tendsto_in_measure_of_tendsto_Lp [hp : fact (1 ≤ p)] {f : ι → Lp E p μ} {g : Lp E p μ} {l : filter ι} (hfg : tendsto f l (𝓝 g)) : tendsto_in_measure μ (λ n, f n) l g := -tendsto_in_measure_of_tendsto_snorm (ennreal.zero_lt_one.trans_le hp.elim).ne.symm +tendsto_in_measure_of_tendsto_snorm (zero_lt_one.trans_le hp.elim).ne.symm (λ n, Lp.ae_strongly_measurable _) (Lp.ae_strongly_measurable _) ((Lp.tendsto_Lp_iff_tendsto_ℒp' _ _).mp hfg) diff --git a/src/measure_theory/function/jacobian.lean b/src/measure_theory/function/jacobian.lean index 999c80fd00745..3be38f43111b4 100644 --- a/src/measure_theory/function/jacobian.lean +++ b/src/measure_theory/function/jacobian.lean @@ -340,7 +340,7 @@ begin by simp only [abs_of_nonneg r0, add_haar_smul, image_add_left, abs_pow, singleton_add, measure_preimage_add] ... ≤ ennreal.of_real (r ^ finrank ℝ E) * (m * μ (closed_ball 0 1)) : - by { rw add_comm, exact ennreal.mul_le_mul le_rfl hε.le } + by { rw add_comm, exact mul_le_mul_left' hε.le _ } ... = m * μ (closed_ball x r) : by { simp only [add_haar_closed_ball' _ _ r0], ring } }, -- covering `s` by closed balls with total measure very close to `μ s`, one deduces that the @@ -366,7 +366,7 @@ begin ... ≤ ∑' (x : t), m * μ (closed_ball x (r x)) : ennreal.tsum_le_tsum (λ x, I x (r x) (ts x.2) (rpos x x.2).le) ... ≤ m * (μ s + a) : - by { rw ennreal.tsum_mul_left, exact ennreal.mul_le_mul le_rfl μt } }, + by { rw ennreal.tsum_mul_left, exact mul_le_mul_left' μt _ } }, -- taking the limit in `a`, one obtains the conclusion have L : tendsto (λ a, (m : ℝ≥0∞) * (μ s + a)) (𝓝[>] 0) (𝓝 (m * (μ s + 0))), { apply tendsto.mono_left _ nhds_within_le_nhds, @@ -587,7 +587,7 @@ begin end ... ≤ ∑' n, (real.to_nnreal (|(A n).det|) + 1 : ℝ≥0) * 0 : begin - refine ennreal.tsum_le_tsum (λ n, ennreal.mul_le_mul le_rfl _), + refine ennreal.tsum_le_tsum (λ n, mul_le_mul_left' _ _), exact le_trans (measure_mono (inter_subset_left _ _)) (le_of_eq hs), end ... = 0 : by simp only [tsum_zero, mul_zero] @@ -642,7 +642,7 @@ begin ... ≤ ε * ∑' n, μ (closed_ball 0 R ∩ t n) : begin rw ennreal.tsum_mul_left, - refine ennreal.mul_le_mul le_rfl (ennreal.tsum_le_tsum (λ n, measure_mono _)), + refine mul_le_mul_left' (ennreal.tsum_le_tsum (λ n, measure_mono _)) _, exact inter_subset_inter_left _ hs, end ... = ε * μ (⋃ n, closed_ball 0 R ∩ t n) : @@ -655,7 +655,7 @@ begin ... ≤ ε * μ (closed_ball 0 R) : begin rw ← inter_Union, - exact ennreal.mul_le_mul le_rfl (measure_mono (inter_subset_left _ _)), + exact mul_le_mul_left' (measure_mono (inter_subset_left _ _)) _, end end @@ -990,7 +990,7 @@ begin { assume t g htg, rcases eq_or_ne (μ t) ∞ with ht|ht, { simp only [ht, εpos.ne', with_top.mul_top, ennreal.coe_eq_zero, le_top, ne.def, - not_false_iff, ennreal.add_top] }, + not_false_iff, _root_.add_top] }, have := h t g (htg.mono_num (min_le_left _ _)), rwa [with_top.coe_sub, ennreal.sub_mul, tsub_le_iff_right] at this, simp only [ht, implies_true_iff, ne.def, not_false_iff] } }, diff --git a/src/measure_theory/function/l2_space.lean b/src/measure_theory/function/l2_space.lean index 04fa719cde5b8..1af86e1ce6aa7 100644 --- a/src/measure_theory/function/l2_space.lean +++ b/src/measure_theory/function/l2_space.lean @@ -34,15 +34,15 @@ variables {α F : Type*} {m : measurable_space α} {μ : measure α} [normed_add lemma mem_ℒp.integrable_sq {f : α → ℝ} (h : mem_ℒp f 2 μ) : integrable (λ x, (f x)^2) μ := by simpa [← mem_ℒp_one_iff_integrable] - using h.norm_rpow ennreal.two_ne_zero ennreal.two_ne_top + using h.norm_rpow 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, + convert (mem_ℒp_norm_rpow_iff hf 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] } + { rw [div_eq_mul_inv, ennreal.mul_inv_cancel two_ne_zero ennreal.two_ne_top] } end lemma mem_ℒp_two_iff_integrable_sq {f : α → ℝ} (hf : ae_strongly_measurable f μ) : @@ -119,11 +119,11 @@ begin have h_two : (2 : ℝ≥0∞).to_real = 2 := by simp, rw [inner_def, integral_inner_eq_sq_snorm, norm_def, ← ennreal.to_real_pow, is_R_or_C.of_real_re, ennreal.to_real_eq_to_real (ennreal.pow_ne_top (Lp.snorm_ne_top f)) _], - { rw [←ennreal.rpow_nat_cast, snorm_eq_snorm' ennreal.two_ne_zero ennreal.two_ne_top, snorm', + { rw [←ennreal.rpow_nat_cast, snorm_eq_snorm' two_ne_zero ennreal.two_ne_top, snorm', ← ennreal.rpow_mul, one_div, h_two], simp, }, { refine (lintegral_rpow_nnnorm_lt_top_of_snorm'_lt_top zero_lt_two _).ne, - rw [← h_two, ← snorm_eq_snorm' ennreal.two_ne_zero ennreal.two_ne_top], + rw [← h_two, ← snorm_eq_snorm' two_ne_zero ennreal.two_ne_top], exact Lp.snorm_lt_top f, }, end diff --git a/src/measure_theory/function/lp_space.lean b/src/measure_theory/function/lp_space.lean index 3f43ae6135b10..f26bde6cd752d 100644 --- a/src/measure_theory/function/lp_space.lean +++ b/src/measure_theory/function/lp_space.lean @@ -1243,8 +1243,7 @@ begin swap, { rw one_div_nonneg, exact ennreal.to_real_nonneg, }, refine lintegral_mono_ae _, filter_upwards [@ennreal.ae_le_ess_sup _ _ μ (λ x, ↑‖φ x‖₊)] with x hx, - refine ennreal.mul_le_mul _ le_rfl, - exact ennreal.rpow_le_rpow hx ennreal.to_real_nonneg, + exact mul_le_mul_right' (ennreal.rpow_le_rpow hx ennreal.to_real_nonneg) _ end ... = ess_sup (λ x, ↑‖φ x‖₊) μ * (∫⁻ x, ↑‖f x‖₊ ^ p.to_real ∂μ) ^ (1 / p.to_real) : begin @@ -1283,12 +1282,12 @@ begin { simp [hp_zero], }, have hq_ne_zero : q ≠ 0, { intro hq_zero, - simp only [hq_zero, hp_zero, one_div, ennreal.inv_zero, ennreal.top_add, + simp only [hq_zero, hp_zero, one_div, ennreal.inv_zero, top_add, ennreal.inv_eq_top] at hpqr, exact hpqr, }, have hr_ne_zero : r ≠ 0, { intro hr_zero, - simp only [hr_zero, hp_zero, one_div, ennreal.inv_zero, ennreal.add_top, + simp only [hr_zero, hp_zero, one_div, ennreal.inv_zero, add_top, ennreal.inv_eq_top] at hpqr, exact hpqr, }, by_cases hq_top : q = ∞, @@ -1744,7 +1743,7 @@ instance [hp : fact (1 ≤ p)] : normed_add_comm_group (Lp E p μ) := rw [snorm_congr_ae (coe_fn_add _ _)], exact snorm_add_le (Lp.ae_strongly_measurable f) (Lp.ae_strongly_measurable g) hp.1, end, - eq_zero_of_map_eq_zero' := λ f, (norm_eq_zero_iff $ ennreal.zero_lt_one.trans_le hp.1).1 } } + eq_zero_of_map_eq_zero' := λ f, (norm_eq_zero_iff $ zero_lt_one.trans_le hp.1).1 } } -- check no diamond is created example [fact (1 ≤ p)] : @@ -1822,7 +1821,7 @@ lemma snorm_ess_sup_indicator_const_le (s : set α) (c : G) : begin by_cases hμ0 : μ = 0, { rw [hμ0, snorm_ess_sup_measure_zero], - exact ennreal.coe_nonneg }, + exact zero_le _ }, { exact (snorm_ess_sup_indicator_le s (λ x, c)).trans (snorm_ess_sup_const c hμ0).le, }, end @@ -2678,7 +2677,7 @@ begin have h_cau' : ∀ (N n m : ℕ), N ≤ n → N ≤ m → snorm' (f n - f m) (p.to_real) μ < B N, { intros N n m hn hm, specialize h_cau N n m hn hm, - rwa snorm_eq_snorm' (ennreal.zero_lt_one.trans_le hp).ne.symm hp_top at h_cau, }, + rwa snorm_eq_snorm' (zero_lt_one.trans_le hp).ne.symm hp_top at h_cau, }, exact ae_tendsto_of_cauchy_snorm' hf hp1 hB h_cau', end @@ -2715,7 +2714,7 @@ lemma mem_ℒp_of_cauchy_tendsto (hp : 1 ≤ p) {f : ℕ → α → E} (hf : ∀ begin refine ⟨h_lim_meas, _⟩, rw ennreal.tendsto_at_top_zero at h_tendsto, - cases (h_tendsto 1 ennreal.zero_lt_one) with N h_tendsto_1, + cases (h_tendsto 1 zero_lt_one) with N h_tendsto_1, specialize h_tendsto_1 N (le_refl N), have h_add : f_lim = f_lim - f N + f N, by abel, rw h_add, diff --git a/src/measure_theory/function/simple_func_dense_lp.lean b/src/measure_theory/function/simple_func_dense_lp.lean index 7185f7c68dd4b..5c0400234c670 100644 --- a/src/measure_theory/function/simple_func_dense_lp.lean +++ b/src/measure_theory/function/simple_func_dense_lp.lean @@ -314,7 +314,7 @@ lemma mem_ℒp_iff {f : α →ₛ E} (hp_pos : p ≠ 0) (hp_ne_top : p ≠ ∞) λ h, mem_ℒp_of_finite_measure_preimage p h⟩ lemma integrable_iff {f : α →ₛ E} : integrable f μ ↔ ∀ y ≠ 0, μ (f ⁻¹' {y}) < ∞ := -mem_ℒp_one_iff_integrable.symm.trans $ mem_ℒp_iff ennreal.zero_lt_one.ne' ennreal.coe_ne_top +mem_ℒp_one_iff_integrable.symm.trans $ mem_ℒp_iff one_ne_zero ennreal.coe_ne_top lemma mem_ℒp_iff_integrable {f : α →ₛ E} (hp_pos : p ≠ 0) (hp_ne_top : p ≠ ∞) : mem_ℒp f p μ ↔ integrable f μ := @@ -849,7 +849,7 @@ lemma Lp.induction [_i : fact (1 ≤ p)] (hp_ne_top : p ≠ ∞) (P : Lp E p μ ∀ f : Lp E p μ, P f := begin refine λ f, (Lp.simple_func.dense_range hp_ne_top).induction_on f h_closed _, - refine Lp.simple_func.induction (lt_of_lt_of_le ennreal.zero_lt_one _i.elim).ne' hp_ne_top _ _, + refine Lp.simple_func.induction (lt_of_lt_of_le zero_lt_one _i.elim).ne' hp_ne_top _ _, { exact λ c s, h_ind c }, { exact λ f g hf hg, h_add hf hg }, end @@ -880,7 +880,7 @@ begin { intros c s hs h, by_cases hc : c = 0, { subst hc, convert h_ind 0 measurable_set.empty (by simp) using 1, ext, simp [const] }, - have hp_pos : p ≠ 0 := (lt_of_lt_of_le ennreal.zero_lt_one _i.elim).ne', + have hp_pos : p ≠ 0 := (lt_of_lt_of_le zero_lt_one _i.elim).ne', exact h_ind c hs (simple_func.measure_lt_top_of_mem_ℒp_indicator hp_pos hp_ne_top hc hs h) }, { intros f g hfg hf hg int_fg, rw [simple_func.coe_add, diff --git a/src/measure_theory/function/uniform_integrable.lean b/src/measure_theory/function/uniform_integrable.lean index 2add4cce0f1c7..6087380fdc8a0 100644 --- a/src/measure_theory/function/uniform_integrable.lean +++ b/src/measure_theory/function/uniform_integrable.lean @@ -484,7 +484,7 @@ begin { simp [indicator_of_not_mem hx] } }, refine le_trans (snorm_mono this) _, rw snorm_indicator_const hs hp hp', - refine ennreal.mul_le_mul (le_of_eq _) le_rfl, + refine mul_le_mul_right' (le_of_eq _) _, rw [← of_real_norm_eq_coe_nnnorm, real.norm_eq_abs, abs_of_nonneg hc], end @@ -508,7 +508,7 @@ begin have hdivp : 0 ≤ 1 / p.to_real, { refine one_div_nonneg.2 _, rw [← ennreal.zero_to_real, ennreal.to_real_le_to_real ennreal.zero_ne_top hp'], - exact le_trans ennreal.zero_lt_one.le hp }, + exact le_trans (zero_le _) hp }, have hpow : 0 < (measure_univ_nnreal μ) ^ (1 / p.to_real) := real.rpow_pos_of_pos (measure_univ_nnreal_pos hμ) _, obtain ⟨δ₁, hδ₁, hsnorm₁⟩ := hui hε', @@ -638,7 +638,7 @@ lemma tendsto_in_measure_iff_tendsto_Lp [is_finite_measure μ] tendsto (λ n, snorm (f n - g) p μ) at_top (𝓝 0) := ⟨λ h, tendsto_Lp_of_tendsto_in_measure μ hp hp' (λ n, (hf n).1) hg h.2 h.1, λ h, ⟨tendsto_in_measure_of_tendsto_snorm - (lt_of_lt_of_le ennreal.zero_lt_one hp).ne.symm + (lt_of_lt_of_le zero_lt_one hp).ne.symm (λ n, (hf n).ae_strongly_measurable) hg.ae_strongly_measurable h, unif_integrable_of_tendsto_Lp μ hp hp' hf hg h⟩⟩ @@ -649,7 +649,7 @@ lemma unif_integrable_of' (hp : 1 ≤ p) (hp' : p ≠ ∞) {f : ι → α → β ∀ i, snorm ({x | C ≤ ‖f i x‖₊}.indicator (f i)) p μ ≤ ennreal.of_real ε) : unif_integrable f p μ := begin - have hpzero := (lt_of_lt_of_le ennreal.zero_lt_one hp).ne.symm, + have hpzero := (lt_of_lt_of_le zero_lt_one hp).ne.symm, by_cases hμ : μ set.univ = 0, { rw measure.measure_univ_eq_zero at hμ, exact hμ.symm ▸ unif_integrable_zero_meas }, @@ -896,7 +896,7 @@ begin begin rw [ennreal.smul_def, ennreal.smul_def, smul_eq_mul, smul_eq_mul], simp_rw ennreal.of_real_coe_nnreal at hℐ, - refine ennreal.mul_le_mul le_rfl (ennreal.rpow_le_rpow (hℐ C).le + refine mul_le_mul' le_rfl (ennreal.rpow_le_rpow (hℐ C).le (one_div_nonneg.2 ennreal.to_real_nonneg)), end ... ≤ snorm ({x | C ≤ ‖f (ℐ C) x‖₊}.indicator (f (ℐ C))) p μ : @@ -942,7 +942,7 @@ lemma uniform_integrable_iff [is_finite_measure μ] (hp : 1 ≤ p) (hp' : p ≠ uniform_integrable f p μ ↔ (∀ i, ae_strongly_measurable (f i) μ) ∧ ∀ ε : ℝ, 0 < ε → ∃ C : ℝ≥0, ∀ i, snorm ({x | C ≤ ‖f i x‖₊}.indicator (f i)) p μ ≤ ennreal.of_real ε := -⟨λ h, ⟨h.1, λ ε, h.spec (lt_of_lt_of_le ennreal.zero_lt_one hp).ne.symm hp'⟩, +⟨λ h, ⟨h.1, λ ε, h.spec (lt_of_lt_of_le zero_lt_one hp).ne.symm hp'⟩, λ h, uniform_integrable_of hp hp' h.1 h.2⟩ diff --git a/src/measure_theory/integral/lebesgue.lean b/src/measure_theory/integral/lebesgue.lean index 8d2c15e95793a..6b78bb4a2bd42 100644 --- a/src/measure_theory/integral/lebesgue.lean +++ b/src/measure_theory/integral/lebesgue.lean @@ -1406,7 +1406,7 @@ begin ... = 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_right (mul_le_mul_left' hs.le _) _ ... ≤ (ε₂ - ε₁) + ε₁ : add_le_add mul_div_le le_rfl ... = ε₂ : tsub_add_cancel_of_le hε₁₂.le, end @@ -1794,7 +1794,7 @@ lemma lintegral_sub_le (f g : α → ℝ≥0∞) (hf : measurable f) : begin rw tsub_le_iff_right, by_cases hfi : ∫⁻ x, f x ∂μ = ∞, - { rw [hfi, ennreal.add_top], + { rw [hfi, add_top], exact le_top }, { rw [← lintegral_add_right _ hf], exact lintegral_mono (λ x, le_tsub_add) } @@ -2685,7 +2685,7 @@ begin { intros g h h_univ h_mea_g h_mea_h h_ind_g h_ind_h, simp [mul_add, *, measurable.mul] }, { intros g h_mea_g h_mono_g h_ind, - have : monotone (λ n a, f a * g n a) := λ m n hmn x, ennreal.mul_le_mul le_rfl (h_mono_g hmn x), + have : monotone (λ n a, f a * g n a) := λ m n hmn x, mul_le_mul_left' (h_mono_g hmn x) _, simp [lintegral_supr, ennreal.mul_supr, h_mf.mul (h_mea_g _), *] } end @@ -2746,7 +2746,7 @@ lemma lintegral_with_density_le_lintegral_mul (μ : measure α) begin rw [← supr_lintegral_measurable_le_eq_lintegral, ← supr_lintegral_measurable_le_eq_lintegral], refine supr₂_le (λ i i_meas, supr_le (λ hi, _)), - have A : f * i ≤ f * g := λ x, ennreal.mul_le_mul le_rfl (hi x), + have A : f * i ≤ f * g := λ x, mul_le_mul_left' (hi x) _, refine le_supr₂_of_le (f * i) (f_meas.mul i_meas) _, exact le_supr_of_le A (le_of_eq (lintegral_with_density_eq_lintegral_mul _ f_meas i_meas)) end @@ -3039,7 +3039,7 @@ lemma exists_absolutely_continuous_is_finite_measure begin obtain ⟨g, gpos, gmeas, hg⟩ : ∃ (g : α → ℝ≥0), (∀ (x : α), 0 < g x) ∧ measurable g ∧ ∫⁻ (x : α), ↑(g x) ∂μ < 1 := - exists_pos_lintegral_lt_of_sigma_finite μ (ennreal.zero_lt_one).ne', + exists_pos_lintegral_lt_of_sigma_finite μ one_ne_zero, refine ⟨μ.with_density (λ x, g x), is_finite_measure_with_density hg.ne_top, _⟩, have : μ = (μ.with_density (λ x, g x)).with_density (λ x, (g x)⁻¹), { have A : (λ (x : α), (g x : ℝ≥0∞)) * (λ (x : α), (↑(g x))⁻¹) = 1, diff --git a/src/measure_theory/integral/mean_inequalities.lean b/src/measure_theory/integral/mean_inequalities.lean index 1e1c756200284..9c902c4f1671a 100644 --- a/src/measure_theory/integral/mean_inequalities.lean +++ b/src/measure_theory/integral/mean_inequalities.lean @@ -116,8 +116,7 @@ begin begin rw lintegral_mul_const' (npf * nqg) _ (by simp [hf_nontop, hg_nontop, hf_nonzero, hg_nonzero, ennreal.mul_eq_top]), - nth_rewrite 1 ←one_mul (npf * nqg), - refine mul_le_mul _ (le_refl (npf * nqg)), + refine mul_le_of_le_one_left' _, have hf1 := lintegral_rpow_fun_mul_inv_snorm_eq_one hpq.pos hf_nonzero hf_nontop, have hg1 := lintegral_rpow_fun_mul_inv_snorm_eq_one hpq.symm.pos hg_nonzero hg_nontop, exact lintegral_mul_le_one_of_lintegral_rpow_eq_one hpq (hf.mul_const _) hf1 hg1, @@ -195,20 +194,20 @@ begin { rw [←ennreal.zero_rpow_of_pos hp0_lt], exact ennreal.rpow_lt_rpow (by simp [zero_lt_one]) hp0_lt, }, have h_rw : (1 / 2) ^ p * (2:ℝ≥0∞) ^ (p - 1) = 1 / 2, - { rw [sub_eq_add_neg, ennreal.rpow_add _ _ ennreal.two_ne_zero ennreal.coe_ne_top, + { rw [sub_eq_add_neg, ennreal.rpow_add _ _ two_ne_zero ennreal.coe_ne_top, ←mul_assoc, ←ennreal.mul_rpow_of_nonneg _ _ hp0, one_div, - ennreal.inv_mul_cancel ennreal.two_ne_zero ennreal.coe_ne_top, ennreal.one_rpow, + ennreal.inv_mul_cancel two_ne_zero ennreal.coe_ne_top, ennreal.one_rpow, one_mul, ennreal.rpow_neg_one], }, rw ←ennreal.mul_le_mul_left (ne_of_lt h_zero_lt_half_rpow).symm _, { rw [mul_add, ← mul_assoc, ← mul_assoc, h_rw, ←ennreal.mul_rpow_of_nonneg _ _ hp0, mul_add], refine ennreal.rpow_arith_mean_le_arith_mean2_rpow (1/2 : ℝ≥0∞) (1/2 : ℝ≥0∞) (f a) (g a) _ hp1, rw [ennreal.div_add_div_same, one_add_one_eq_two, - ennreal.div_self ennreal.two_ne_zero ennreal.coe_ne_top], }, + ennreal.div_self two_ne_zero ennreal.coe_ne_top], }, { rw ← lt_top_iff_ne_top, refine ennreal.rpow_lt_top_of_nonneg hp0 _, rw [one_div, ennreal.inv_ne_top], - exact ennreal.two_ne_zero, }, + exact two_ne_zero, }, end ... < ⊤ : begin diff --git a/src/measure_theory/integral/vitali_caratheodory.lean b/src/measure_theory/integral/vitali_caratheodory.lean index fd375ddc1f4c5..e5e59cfb7fbdd 100644 --- a/src/measure_theory/integral/vitali_caratheodory.lean +++ b/src/measure_theory/integral/vitali_caratheodory.lean @@ -95,8 +95,7 @@ begin induction f using measure_theory.simple_func.induction with c s hs f₁ f₂ H h₁ h₂ generalizing ε, { let f := simple_func.piecewise s hs (simple_func.const α c) (simple_func.const α 0), by_cases h : ∫⁻ x, f x ∂μ = ⊤, - { refine ⟨λ x, c, λ x, _, lower_semicontinuous_const, - by simp only [ennreal.top_add, le_top, h]⟩, + { refine ⟨λ x, c, λ x, _, lower_semicontinuous_const, by simp only [_root_.top_add, le_top, h]⟩, simp only [simple_func.coe_const, simple_func.const_zero, simple_func.coe_zero, set.piecewise_eq_indicator, simple_func.coe_piecewise], exact set.indicator_le_self _ _ _ }, @@ -125,7 +124,7 @@ begin lintegral_const, ennreal.coe_indicator, set.univ_inter, measurable_set.univ, simple_func.const_zero, lintegral_indicator, simple_func.coe_zero, set.piecewise_eq_indicator, simple_func.coe_piecewise, restrict_apply], - calc (c : ℝ≥0∞) * μ u ≤ c * (μ s + ε / c) : ennreal.mul_le_mul le_rfl μu.le + calc (c : ℝ≥0∞) * μ u ≤ c * (μ s + ε / c) : mul_le_mul_left' μu.le _ ... = c * μ s + ε : begin simp_rw [mul_add], @@ -338,7 +337,7 @@ begin lintegral_const, ennreal.coe_indicator, set.univ_inter, measurable_set.univ, simple_func.const_zero, lintegral_indicator, simple_func.coe_zero, set.piecewise_eq_indicator, simple_func.coe_piecewise, restrict_apply], - calc (c : ℝ≥0∞) * μ s ≤ c * (μ F + ε / c) : ennreal.mul_le_mul le_rfl μF.le + calc (c : ℝ≥0∞) * μ s ≤ c * (μ F + ε / c) : mul_le_mul_left' μF.le _ ... = c * μ F + ε : begin simp_rw [mul_add], diff --git a/src/measure_theory/measure/doubling.lean b/src/measure_theory/measure/doubling.lean index 2fbc89fc639ac..26608590ed104 100644 --- a/src/measure_theory/measure/doubling.lean +++ b/src/measure_theory/measure/doubling.lean @@ -104,7 +104,7 @@ begin refine le_mul_of_one_le_of_le _ le_rfl, apply ennreal.one_le_coe_iff.2 (le_max_right _ _) }, { apply (hR ⟨rpos, hr⟩ x t ht.2).trans _, - exact ennreal.mul_le_mul (ennreal.coe_le_coe.2 (le_max_left _ _)) le_rfl } + exact mul_le_mul_right' (ennreal.coe_le_coe.2 (le_max_left _ _)) _ } end lemma eventually_measure_le_scaling_constant_mul (K : ℝ) : @@ -113,7 +113,7 @@ lemma eventually_measure_le_scaling_constant_mul (K : ℝ) : begin filter_upwards [classical.some_spec (exists_eventually_forall_measure_closed_ball_le_mul μ K)] with r hr x, - exact (hr x K le_rfl).trans (ennreal.mul_le_mul (ennreal.coe_le_coe.2 (le_max_left _ _)) le_rfl) + exact (hr x K le_rfl).trans (mul_le_mul_right' (ennreal.coe_le_coe.2 (le_max_left _ _)) _) end lemma eventually_measure_le_scaling_constant_mul' (K : ℝ) (hK : 0 < K) : diff --git a/src/measure_theory/measure/haar.lean b/src/measure_theory/measure/haar.lean index 4e38aba3cf793..c988a723a3a6c 100644 --- a/src/measure_theory/measure/haar.lean +++ b/src/measure_theory/measure/haar.lean @@ -498,7 +498,7 @@ by simpa only [ennreal.coe_eq_coe, ←nnreal.coe_eq, haar_content_apply] lemma haar_content_outer_measure_self_pos {K₀ : positive_compacts G} : 0 < (haar_content K₀).outer_measure K₀ := begin - apply ennreal.zero_lt_one.trans_le, + refine zero_lt_one.trans_le _, rw [content.outer_measure_eq_infi], refine le_infi₂ (λ U hU, le_infi $ λ hK₀, le_trans _ $ le_supr₂ K₀.to_compacts hK₀), exact haar_content_self.ge, diff --git a/src/measure_theory/measure/haar_lebesgue.lean b/src/measure_theory/measure/haar_lebesgue.lean index 62e0626d6fb24..cfc90c4456988 100644 --- a/src/measure_theory/measure/haar_lebesgue.lean +++ b/src/measure_theory/measure/haar_lebesgue.lean @@ -637,7 +637,7 @@ begin (eventually_of_forall (λ b, zero_le _)), filter_upwards [self_mem_nhds_within], rintros r (rpos : 0 < r), - apply ennreal.mul_le_mul (measure_mono (inter_subset_inter_right _ _)) le_rfl, + apply mul_le_mul_right' (measure_mono (inter_subset_inter_right _ _)) _, assume y hy, have : y - x ∈ r • closed_ball (0 : E) 1, { apply smul_set_mono t_bound, @@ -757,7 +757,7 @@ begin add_le_add le_rfl (measure_mono (inter_subset_right _ _)), calc μ (s ∩ ({x} + r • t)) / μ ({x} + r • t) ≤ (μ (s ∩ ({x} + r • (t ∩ closed_ball 0 n))) + μ ({x} + r • (t \ closed_ball 0 n))) / - μ ({x} + r • t) : ennreal.mul_le_mul I le_rfl + μ ({x} + r • t) : mul_le_mul_right' I _ ... < ε / 2 + ε / 2 : begin rw ennreal.add_div, @@ -836,8 +836,7 @@ begin { apply tendsto_add_haar_inter_smul_one_of_density_one_aux μ _ (measurable_set_to_measurable _ _) _ _ t ht h't h''t, apply tendsto_of_tendsto_of_tendsto_of_le_of_le' h tendsto_const_nhds, - { apply eventually_of_forall (λ r, _), - apply ennreal.mul_le_mul _ le_rfl, + { refine eventually_of_forall (λ r, mul_le_mul_right' _ _), exact measure_mono (inter_subset_inter_left _ (subset_to_measurable _ _)) }, { filter_upwards [self_mem_nhds_within], rintros r (rpos : 0 < r), @@ -864,7 +863,7 @@ begin exists_subset_measure_lt_top ht h't.bot_lt, filter_upwards [(tendsto_order.1 (tendsto_add_haar_inter_smul_one_of_density_one μ s x h t' - t'_meas t'pos.ne' t'top.ne)).1 0 ennreal.zero_lt_one], + t'_meas t'pos.ne' t'top.ne)).1 0 zero_lt_one], assume r hr, have : μ (s ∩ ({x} + r • t')) ≠ 0 := λ h', by simpa only [ennreal.not_lt_zero, ennreal.zero_div, h'] using hr, diff --git a/src/measure_theory/measure/hausdorff.lean b/src/measure_theory/measure/hausdorff.lean index f25584df4b9b0..2b8222aecfe00 100644 --- a/src/measure_theory/measure/hausdorff.lean +++ b/src/measure_theory/measure/hausdorff.lean @@ -184,7 +184,7 @@ begin `μ (s ∩ t) + μ (⋃ n, S n) ≤ μ s`. We can't pass to the limit because `μ` is only an outer measure. -/ by_cases htop : μ (s \ t) = ∞, - { rw [htop, ennreal.add_top, ← htop], + { rw [htop, add_top, ← htop], exact μ.mono (diff_subset _ _) }, suffices : μ (⋃ n, S n) ≤ ⨆ n, μ (S n), calc μ (s ∩ t) + μ (s \ t) = μ (s ∩ t) + μ (⋃ n, S n) : @@ -198,7 +198,7 @@ begin and the second term tends to zero, see `outer_measure.Union_nat_of_monotone_of_tsum_ne_top` for details. -/ have : ∀ n, S n ⊆ S (n + 1), from λ n x hx, - ⟨hx.1, le_trans (ennreal.inv_le_inv.2 $ ennreal.coe_nat_le_coe_nat.2 n.le_succ) hx.2⟩, + ⟨hx.1, le_trans (ennreal.inv_le_inv.2 $ nat.cast_le.2 n.le_succ) hx.2⟩, refine (μ.Union_nat_of_monotone_of_tsum_ne_top this _).le, clear this, /- While the sets `S (k + 1) \ S k` are not pairwise metric separated, the sets in each subsequence `S (2 * k + 1) \ S (2 * k)` and `S (2 * k + 2) \ S (2 * k)` are metric separated, @@ -217,7 +217,7 @@ begin (λ h, (this j i h).symm.mono (λ x hx, ⟨hx.1.1, hx.2⟩) (inter_subset_left _ _)), intros i j hj, have A : ((↑(2 * j + r))⁻¹ : ℝ≥0∞) < (↑(2 * i + 1 + r))⁻¹, - by { rw [ennreal.inv_lt_inv, ennreal.coe_nat_lt_coe_nat], linarith }, + by { rw [ennreal.inv_lt_inv, nat.cast_lt], linarith }, refine ⟨(↑(2 * i + 1 + r))⁻¹ - (↑(2 * j + r))⁻¹, by simpa using A, λ x hx y hy, _⟩, have : inf_edist y t < (↑(2 * j + r))⁻¹, from not_le.1 (λ hle, hy.2 ⟨hy.1, hle⟩), rcases inf_edist_lt_iff.mp this with ⟨z, hzt, hyz⟩, @@ -344,7 +344,7 @@ lemma mk_metric_mono_smul {m₁ m₂ : ℝ≥0∞ → ℝ≥0∞} {c : ℝ≥0 (mk_metric m₁ : outer_measure X) ≤ c • mk_metric m₂ := begin classical, - rcases (mem_nhds_within_Ici_iff_exists_Ico_subset' ennreal.zero_lt_one).1 hle with ⟨r, hr0, hr⟩, + rcases (mem_nhds_within_Ici_iff_exists_Ico_subset' zero_lt_one).1 hle with ⟨r, hr0, hr⟩, refine λ s, le_of_tendsto_of_tendsto (mk_metric'.tendsto_pre _ s) (ennreal.tendsto.const_mul (mk_metric'.tendsto_pre _ s) (or.inr hc)) (mem_of_superset (Ioo_mem_nhds_within_Ioi ⟨le_rfl, hr0⟩) (λ r' hr', _)), @@ -362,7 +362,7 @@ end `mk_metric m₁ hm₁ ≤ mk_metric m₂ hm₂`-/ lemma mk_metric_mono {m₁ m₂ : ℝ≥0∞ → ℝ≥0∞} (hle : m₁ ≤ᶠ[𝓝[≥] 0] m₂) : (mk_metric m₁ : outer_measure X) ≤ mk_metric m₂ := -by { convert mk_metric_mono_smul ennreal.one_ne_top ennreal.zero_lt_one.ne' _; simp * } +by { convert mk_metric_mono_smul ennreal.one_ne_top one_ne_zero _; simp * } lemma isometry_comap_mk_metric (m : ℝ≥0∞ → ℝ≥0∞) {f : X → Y} (hf : isometry f) (H : monotone m ∨ surjective f) : @@ -469,7 +469,7 @@ end `mk_metric m₁ hm₁ ≤ mk_metric m₂ hm₂`-/ lemma mk_metric_mono {m₁ m₂ : ℝ≥0∞ → ℝ≥0∞} (hle : m₁ ≤ᶠ[𝓝[≥] 0] m₂) : (mk_metric m₁ : measure X) ≤ mk_metric m₂ := -by { convert mk_metric_mono_smul ennreal.one_ne_top ennreal.zero_lt_one.ne' _; simp * } +by { convert mk_metric_mono_smul ennreal.one_ne_top one_ne_zero _; simp * } /-- A formula for `measure_theory.measure.mk_metric`. -/ lemma mk_metric_apply (m : ℝ≥0∞ → ℝ≥0∞) (s : set X) : @@ -603,10 +603,9 @@ begin (or.inr $ mt ennreal.coe_eq_zero.1 hc)], rcases eq_or_ne r 0 with rfl|hr₀, { rcases lt_or_le 0 d₂ with h₂|h₂, - { simp only [h₂, ennreal.zero_rpow_of_pos, zero_le', ennreal.coe_nonneg, ennreal.zero_div, - ennreal.coe_zero] }, - { simp only [h.trans_le h₂, ennreal.div_top, zero_le', ennreal.coe_nonneg, - ennreal.zero_rpow_of_neg, ennreal.coe_zero] } }, + { simp only [h₂, ennreal.zero_rpow_of_pos, zero_le, ennreal.zero_div, ennreal.coe_zero] }, + { simp only [h.trans_le h₂, ennreal.div_top, zero_le, ennreal.zero_rpow_of_neg, + ennreal.coe_zero] } }, { have : (r : ℝ≥0∞) ≠ 0, by simpa only [ennreal.coe_eq_zero, ne.def] using hr₀, rw [← ennreal.rpow_sub _ _ this ennreal.coe_ne_top], refine (ennreal.rpow_lt_rpow hrc (sub_pos.2 h)).le.trans _, @@ -648,7 +647,7 @@ begin suffices : (1 : ℝ≥0∞) ≤ ⨅ (t : ℕ → set X) (hts : {x} ⊆ ⋃ n, t n) (ht : ∀ n, diam (t n) ≤ 1), ∑' n, ⨆ (h : (t n).nonempty), (diam (t n)) ^ (0 : ℝ), { apply le_trans this _, - convert le_supr₂ (1 : ℝ≥0∞) (ennreal.zero_lt_one), + convert le_supr₂ (1 : ℝ≥0∞) zero_lt_one, refl }, simp only [ennreal.rpow_zero, le_infi_iff], assume t hst h't, diff --git a/src/measure_theory/measure/lebesgue.lean b/src/measure_theory/measure/lebesgue.lean index fde7b7ecdfdbe..e9cd6f6834cd9 100644 --- a/src/measure_theory/measure/lebesgue.lean +++ b/src/measure_theory/measure/lebesgue.lean @@ -93,7 +93,7 @@ by rw [closed_ball_eq_Icc, volume_Icc, ← sub_add, add_sub_cancel', two_mul] volume (emetric.ball a r) = 2 * r := begin rcases eq_or_ne r ∞ with rfl|hr, - { rw [metric.emetric_ball_top, volume_univ, two_mul, ennreal.top_add] }, + { rw [metric.emetric_ball_top, volume_univ, two_mul, _root_.top_add] }, { lift r to ℝ≥0 using hr, rw [metric.emetric_ball_nnreal, volume_ball, two_mul, ← nnreal.coe_add, ennreal.of_real_coe_nnreal, ennreal.coe_add, two_mul] } @@ -103,7 +103,7 @@ end volume (emetric.closed_ball a r) = 2 * r := begin rcases eq_or_ne r ∞ with rfl|hr, - { rw [emetric.closed_ball_top, volume_univ, two_mul, ennreal.top_add] }, + { rw [emetric.closed_ball_top, volume_univ, two_mul, _root_.top_add] }, { lift r to ℝ≥0 using hr, rw [metric.emetric_closed_ball_nnreal, volume_closed_ball, two_mul, ← nnreal.coe_add, ennreal.of_real_coe_nnreal, ennreal.coe_add, two_mul] } diff --git a/src/measure_theory/measure/measure_space.lean b/src/measure_theory/measure/measure_space.lean index 5a7873f4399c1..86424d7edcee0 100644 --- a/src/measure_theory/measure/measure_space.lean +++ b/src/measure_theory/measure/measure_space.lean @@ -3121,7 +3121,8 @@ lemma countable_meas_pos_of_disjoint_of_meas_Union_ne_top {ι : Type*} [measurab set.countable {i : ι | 0 < μ (As i)} := begin set posmeas := {i : ι | 0 < μ (As i)} with posmeas_def, - rcases exists_seq_strict_anti_tendsto' ennreal.zero_lt_one with ⟨as, ⟨as_decr, ⟨as_mem, as_lim⟩⟩⟩, + rcases exists_seq_strict_anti_tendsto' (zero_lt_one : (0 : ℝ≥0∞) < 1) + with ⟨as, as_decr, as_mem, as_lim⟩, set fairmeas := λ (n : ℕ) , {i : ι | as n ≤ μ (As i)} with fairmeas_def, have countable_union : posmeas = (⋃ n, fairmeas n) , { have fairmeas_eq : ∀ n, fairmeas n = (λ i, μ (As i)) ⁻¹' Ici (as n), diff --git a/src/measure_theory/measure/regular.lean b/src/measure_theory/measure/regular.lean index b5434ca92d550..a025ef6c8720b 100644 --- a/src/measure_theory/measure/regular.lean +++ b/src/measure_theory/measure/regular.lean @@ -259,10 +259,9 @@ lemma _root_.set.exists_is_open_le_add (A : set α) (μ : measure α) [outer_reg {ε : ℝ≥0∞} (hε : ε ≠ 0) : ∃ U ⊇ A, is_open U ∧ μ U ≤ μ A + ε := begin - rcases le_or_lt ∞ (μ A) with H|H, - { exact ⟨univ, subset_univ _, is_open_univ, - by simp only [top_le_iff.mp H, ennreal.top_add, le_top]⟩ }, - { rcases A.exists_is_open_lt_add H.ne hε with ⟨U, AU, U_open, hU⟩, + rcases eq_or_ne (μ A) ∞ with H|H, + { exact ⟨univ, subset_univ _, is_open_univ, by simp only [H, _root_.top_add, le_top]⟩ }, + { rcases A.exists_is_open_lt_add H hε with ⟨U, AU, U_open, hU⟩, exact ⟨U, AU, U_open, hU.le⟩ } end diff --git a/src/order/filter/ennreal.lean b/src/order/filter/ennreal.lean index 580c082a1d442..affc7d7eae17d 100644 --- a/src/order/filter/ennreal.lean +++ b/src/order/filter/ennreal.lean @@ -100,8 +100,7 @@ lemma limsup_mul_le [countable_Inter_filter f] (u v : α → ℝ≥0∞) : calc f.limsup (u * v) ≤ f.limsup (λ x, (f.limsup u) * v x) : begin refine limsup_le_limsup _ _, - { filter_upwards [@eventually_le_limsup _ f _ u] with x hx, - exact ennreal.mul_le_mul hx le_rfl, }, + { filter_upwards [@eventually_le_limsup _ f _ u] with x hx using mul_le_mul_right' hx _ }, { is_bounded_default, }, end ... = f.limsup u * f.limsup v : limsup_const_mul diff --git a/src/probability/integration.lean b/src/probability/integration.lean index 20776ca4089ee..d889e879b0b2e 100644 --- a/src/probability/integration.lean +++ b/src/probability/integration.lean @@ -67,7 +67,7 @@ begin rw [lintegral_supr h_measM_f h_mono_f, lintegral_supr, ennreal.supr_mul], { simp_rw [← h_ind_f] }, { exact λ n, h_mul_indicator _ (h_measM_f n) }, - { exact λ m n h_le a, ennreal.mul_le_mul (h_mono_f h_le a) le_rfl, }, }, + { exact λ m n h_le a, mul_le_mul_right' (h_mono_f h_le a) _, }, }, end /-- If `f` and `g` are independent random variables with values in `ℝ≥0∞`, @@ -101,7 +101,7 @@ begin rw [lintegral_supr, lintegral_supr h_measM_f' h_mono_f', ennreal.mul_supr], { simp_rw [← h_ind_f'], }, { exact λ n, h_measM_f.mul (h_measM_f' n), }, - { exact λ n m (h_le : n ≤ m) a, ennreal.mul_le_mul le_rfl (h_mono_f' h_le a), }, } + { exact λ n m (h_le : n ≤ m) a, mul_le_mul_left' (h_mono_f' h_le a) _, }, } end /-- If `f` and `g` are independent random variables with values in `ℝ≥0∞`, diff --git a/src/probability/kernel/basic.lean b/src/probability/kernel/basic.lean index 6a51e139debc6..22eb09bd1b147 100644 --- a/src/probability/kernel/basic.lean +++ b/src/probability/kernel/basic.lean @@ -663,7 +663,7 @@ begin ≤ ∫⁻ b in set.univ, B ∂(κ a) : lintegral_mono (hf_B a) ... = B * κ a set.univ : by simp only [measure.restrict_univ, lintegral_const] ... ≤ B * is_finite_kernel.bound κ : - ennreal.mul_le_mul le_rfl (measure_le_bound κ a set.univ), + mul_le_mul_left' (measure_le_bound κ a set.univ) _, end⟩⟩, }, { rw with_density_of_not_measurable _ hf, apply_instance, }, diff --git a/src/probability/martingale/upcrossing.lean b/src/probability/martingale/upcrossing.lean index c2d3cb4769399..df61f68c204c9 100644 --- a/src/probability/martingale/upcrossing.lean +++ b/src/probability/martingale/upcrossing.lean @@ -873,10 +873,10 @@ begin simp_rw [this, upcrossings, supr_le_iff], split; rintro ⟨k, hk⟩, { obtain ⟨m, hm⟩ := exists_nat_ge k, - refine ⟨m, λ N, ennreal.coe_nat_le_coe_nat.1 ((hk N).trans _)⟩, + refine ⟨m, λ N, nat.cast_le.1 ((hk N).trans _)⟩, rwa [← ennreal.coe_nat, ennreal.coe_le_coe] }, { refine ⟨k, λ N, _⟩, - simp only [ennreal.coe_nat, ennreal.coe_nat_le_coe_nat, hk N] } + simp only [ennreal.coe_nat, nat.cast_le, hk N] } end /-- A variant of Doob's upcrossing estimate obtained by taking the supremum on both sides. -/ @@ -905,7 +905,7 @@ begin { exact λ n, measurable_from_top.comp_ae_measurable (hf.adapted.measurable_upcrossings_before hab).ae_measurable }, { refine eventually_of_forall (λ ω N M hNM, _), - rw ennreal.coe_nat_le_coe_nat, + rw nat.cast_le, exact upcrossings_before_mono hab hNM ω } }, { rw [not_lt, ← sub_nonpos] at hab, rw [ennreal.of_real_of_nonpos hab, zero_mul], diff --git a/src/probability/variance.lean b/src/probability/variance.lean index 2e539b0e9810e..ce7e4589021ff 100644 --- a/src/probability/variance.lean +++ b/src/probability/variance.lean @@ -55,7 +55,7 @@ lemma _root_.measure_theory.mem_ℒp.evariance_lt_top [is_finite_measure μ] (hX evariance X μ < ∞ := begin have := ennreal.pow_lt_top (hX.sub $ mem_ℒp_const $ μ[X]).2 2, - rw [snorm_eq_lintegral_rpow_nnnorm ennreal.two_ne_zero ennreal.two_ne_top, + rw [snorm_eq_lintegral_rpow_nnnorm two_ne_zero ennreal.two_ne_top, ← ennreal.rpow_two] at this, simp only [pi.sub_apply, ennreal.to_real_bit0, ennreal.one_to_real, one_div] at this, rw [← ennreal.rpow_mul, inv_mul_cancel (two_ne_zero : (2 : ℝ) ≠ 0), ennreal.rpow_one] at this, @@ -71,7 +71,7 @@ begin rw [← ne.def, ← lt_top_iff_ne_top] at h, have : mem_ℒp (λ ω, X ω - μ[X]) 2 μ, { refine ⟨hXm.sub ae_strongly_measurable_const, _⟩, - rw snorm_eq_lintegral_rpow_nnnorm ennreal.two_ne_zero ennreal.two_ne_top, + rw snorm_eq_lintegral_rpow_nnnorm two_ne_zero ennreal.two_ne_top, simp only [ennreal.to_real_bit0, ennreal.one_to_real, ennreal.rpow_two, ne.def], exact ennreal.rpow_lt_top_of_nonneg (by simp) h.ne }, refine hX _, @@ -117,7 +117,7 @@ begin simp_rw [hXint, sub_zero], { refl }, { exact integral_nonneg (λ ω, pow_two_nonneg _) }, - { convert hX.integrable_norm_rpow ennreal.two_ne_zero ennreal.two_ne_top, + { convert hX.integrable_norm_rpow two_ne_zero ennreal.two_ne_top, ext ω, simp only [pi.sub_apply, real.norm_eq_abs, ennreal.to_real_bit0, ennreal.one_to_real, real.rpow_two, pow_bit0_abs] }, @@ -133,7 +133,7 @@ begin { refl }, { exact integral_nonneg (λ ω, pow_two_nonneg _) }, { convert (hX.sub $ mem_ℒp_const (μ[X])).integrable_norm_rpow - ennreal.two_ne_zero ennreal.two_ne_top, + two_ne_zero ennreal.two_ne_top, ext ω, simp only [pi.sub_apply, real.norm_eq_abs, ennreal.to_real_bit0, ennreal.one_to_real, real.rpow_two, pow_bit0_abs] }, @@ -267,7 +267,7 @@ begin refine ⟨_, ennreal.of_real_ne_top⟩, rw [mem_ℒp, not_and] at hℒ, specialize hℒ hX, - simp only [snorm_eq_lintegral_rpow_nnnorm ennreal.two_ne_zero ennreal.two_ne_top, not_lt, + simp only [snorm_eq_lintegral_rpow_nnnorm two_ne_zero ennreal.two_ne_top, not_lt, top_le_iff, ennreal.to_real_bit0, ennreal.one_to_real, ennreal.rpow_two, one_div, ennreal.rpow_eq_top_iff, inv_lt_zero, inv_pos, zero_lt_bit0, zero_lt_one, and_true, or_iff_not_imp_left, not_and_distrib] at hℒ, @@ -281,11 +281,11 @@ theorem meas_ge_le_evariance_div_sq {X : Ω → ℝ} begin have A : (c : ℝ≥0∞) ≠ 0, { rwa [ne.def, ennreal.coe_eq_zero] }, have B : ae_strongly_measurable (λ (ω : Ω), 𝔼[X]) ℙ := ae_strongly_measurable_const, - convert meas_ge_le_mul_pow_snorm ℙ ennreal.two_ne_zero ennreal.two_ne_top (hX.sub B) A, + convert meas_ge_le_mul_pow_snorm ℙ two_ne_zero ennreal.two_ne_top (hX.sub B) A, { ext ω, simp only [pi.sub_apply, ennreal.coe_le_coe, ← real.norm_eq_abs, ← coe_nnnorm, nnreal.coe_le_coe, ennreal.of_real_coe_nnreal] }, - { rw snorm_eq_lintegral_rpow_nnnorm ennreal.two_ne_zero ennreal.two_ne_top, + { rw snorm_eq_lintegral_rpow_nnnorm two_ne_zero ennreal.two_ne_top, simp only [ennreal.to_real_bit0, ennreal.one_to_real, pi.sub_apply, one_div], rw [div_eq_mul_inv, ennreal.inv_pow, mul_comm, ennreal.rpow_two], congr, diff --git a/src/topology/instances/ennreal.lean b/src/topology/instances/ennreal.lean index 49c3479737fbd..165fcb8f6e80f 100644 --- a/src/topology/instances/ennreal.lean +++ b/src/topology/instances/ennreal.lean @@ -161,7 +161,7 @@ tendsto_nhds_top_iff_nat.2 h lemma tendsto_nat_nhds_top : tendsto (λ n : ℕ, ↑n) at_top (𝓝 ∞) := tendsto_nhds_top $ λ n, mem_at_top_sets.2 - ⟨n+1, λ m hm, ennreal.coe_nat_lt_coe_nat.2 $ nat.lt_of_succ_le hm⟩ + ⟨n + 1, λ m hm, mem_set_of.2 $ nat.cast_lt.2 $ nat.lt_of_succ_le hm⟩ @[simp, norm_cast] lemma tendsto_coe_nhds_top {f : α → ℝ≥0} {l : filter α} : tendsto (λ x, (f x : ℝ≥0∞)) l (𝓝 ∞) ↔ tendsto f l at_top := @@ -171,7 +171,7 @@ by rw [tendsto_nhds_top_iff_nnreal, at_top_basis_Ioi.tendsto_right_iff]; lemma tendsto_of_real_at_top : tendsto ennreal.of_real at_top (𝓝 ∞) := tendsto_coe_nhds_top.2 tendsto_real_to_nnreal_at_top -lemma nhds_zero : 𝓝 (0 : ℝ≥0∞) = ⨅a ≠ 0, 𝓟 (Iio a) := +lemma nhds_zero : 𝓝 (0 : ℝ≥0∞) = ⨅ a ≠ 0, 𝓟 (Iio a) := nhds_bot_order.trans $ by simp [bot_lt_iff_ne_bot, Iio] lemma nhds_zero_basis : (𝓝 (0 : ℝ≥0∞)).has_basis (λ a : ℝ≥0∞, 0 < a) (λ a, Iio a) := nhds_bot_basis @@ -449,7 +449,7 @@ begin have : tendsto (* x) (𝓝[<] 1) (𝓝 (1 * x)) := (ennreal.continuous_at_mul_const (or.inr one_ne_zero)).mono_left inf_le_left, rw one_mul at this, - haveI : (𝓝[<] (1 : ℝ≥0∞)).ne_bot := nhds_within_Iio_self_ne_bot' ⟨0, ennreal.zero_lt_one⟩, + haveI : (𝓝[<] (1 : ℝ≥0∞)).ne_bot := nhds_within_Iio_self_ne_bot' ⟨0, zero_lt_one⟩, exact le_of_tendsto this (eventually_nhds_within_iff.2 $ eventually_of_forall h) end diff --git a/src/topology/metric_space/emetric_paracompact.lean b/src/topology/metric_space/emetric_paracompact.lean index 554e959682fde..1a1604f119859 100644 --- a/src/topology/metric_space/emetric_paracompact.lean +++ b/src/topology/metric_space/emetric_paracompact.lean @@ -39,7 +39,7 @@ begin have pow_pos : ∀ k : ℕ, (0 : ℝ≥0∞) < 2⁻¹ ^ k, from λ k, ennreal.pow_pos (ennreal.inv_pos.2 ennreal.two_ne_top) _, have hpow_le : ∀ {m n : ℕ}, m ≤ n → (2⁻¹ : ℝ≥0∞) ^ n ≤ 2⁻¹ ^ m, - from λ m n h, ennreal.pow_le_pow_of_le_one (ennreal.inv_le_one.2 ennreal.one_lt_two.le) h, + from λ m n h, pow_le_pow_of_le_one' (ennreal.inv_le_one.2 ennreal.one_lt_two.le) h, have h2pow : ∀ n : ℕ, 2 * (2⁻¹ : ℝ≥0∞) ^ (n + 1) = 2⁻¹ ^ n, by { intro n, simp [pow_succ, ← mul_assoc, ennreal.mul_inv_cancel] }, -- Consider an open covering `S : set (set α)` @@ -99,9 +99,9 @@ begin rintro ⟨y, rfl, hsub, -, hyx⟩, refine hsub (lt_of_lt_of_le hyx _), calc 2⁻¹ ^ n = 1 * 2⁻¹ ^ n : (one_mul _).symm - ... ≤ 3 * 2⁻¹ ^ n : ennreal.mul_le_mul _ le_rfl, + ... ≤ 3 * 2⁻¹ ^ n : mul_le_mul_right' _ _, -- TODO: use `norm_num` - have : ((1 : ℕ) : ℝ≥0∞) ≤ (3 : ℕ), from ennreal.coe_nat_le_coe_nat.2 (by norm_num1), + have : ((1 : ℕ) : ℝ≥0∞) ≤ (3 : ℕ), from nat.cast_le.2 (by norm_num1), exact_mod_cast this }, -- Let us show the rest of the properties. Since the definition expects a family indexed -- by a single parameter, we use `ℕ × ι` as the domain. @@ -148,7 +148,7 @@ begin by apply_rules [ennreal.add_lt_add] ... = 2 * (2⁻¹ ^ m + 2⁻¹ ^ (n + k + 1)) : by simp only [two_mul, add_comm] ... ≤ 2 * (2⁻¹ ^ m + 2⁻¹ ^ (m + 1)) : - ennreal.mul_le_mul le_rfl $ add_le_add le_rfl $ hpow_le (add_le_add hm le_rfl) + mul_le_mul' le_rfl $ add_le_add le_rfl $ hpow_le (add_le_add hm le_rfl) ... = 3 * 2⁻¹ ^ m : by rw [mul_add, h2pow, bit1, add_mul, one_mul] }, -- Finally, we glue `Hgt` and `Hle` have : (⋃ (m ≤ n + k) (i ∈ {i : ι | (D m i ∩ B).nonempty}), {(m, i)}).finite := diff --git a/src/topology/metric_space/emetric_space.lean b/src/topology/metric_space/emetric_space.lean index acbfce8f93ca4..64cde1bbe17d1 100644 --- a/src/topology/metric_space/emetric_space.lean +++ b/src/topology/metric_space/emetric_space.lean @@ -181,7 +181,7 @@ theorem uniformity_basis_edist : (λ r hr p hp, ⟨min r p, lt_min hr hp, λ x hx, lt_of_lt_of_le hx (min_le_left _ _), λ x hx, lt_of_lt_of_le hx (min_le_right _ _)⟩) - ⟨1, ennreal.zero_lt_one⟩ + ⟨1, zero_lt_one⟩ /-- Characterization of the elements of the uniformity in terms of the extended distance -/ theorem mem_uniformity_edist {s : set (α×α)} : @@ -645,7 +645,7 @@ is_open_iff.2 $ λ y, exists_ball_subset_ball theorem is_closed_ball_top : is_closed (ball x ⊤) := is_open_compl_iff.1 $ is_open_iff.2 $ λ y hy, ⟨⊤, ennreal.coe_lt_top, - (ball_disjoint $ by { rw ennreal.top_add, exact le_of_not_lt hy }).subset_compl_right⟩ + (ball_disjoint $ by { rw top_add, exact le_of_not_lt hy }).subset_compl_right⟩ theorem ball_mem_nhds (x : α) {ε : ℝ≥0∞} (ε0 : 0 < ε) : ball x ε ∈ 𝓝 x := is_open_ball.mem_nhds (mem_ball_self ε0) diff --git a/src/topology/metric_space/hausdorff_dimension.lean b/src/topology/metric_space/hausdorff_dimension.lean index 9267ca31287b9..da666adf355ce 100644 --- a/src/topology/metric_space/hausdorff_dimension.lean +++ b/src/topology/metric_space/hausdorff_dimension.lean @@ -532,4 +532,4 @@ in `F`. -/ lemma cont_diff.dense_compl_range_of_finrank_lt_finrank [finite_dimensional ℝ F] {f : E → F} (h : cont_diff ℝ 1 f) (hEF : finrank ℝ E < finrank ℝ F) : dense (range f)ᶜ := -dense_compl_of_dimH_lt_finrank $ h.dimH_range_le.trans_lt $ ennreal.coe_nat_lt_coe_nat.2 hEF +dense_compl_of_dimH_lt_finrank $ h.dimH_range_le.trans_lt $ nat.cast_lt.2 hEF diff --git a/src/topology/metric_space/hausdorff_distance.lean b/src/topology/metric_space/hausdorff_distance.lean index 2a93ac1486ed0..3c498da5f248b 100644 --- a/src/topology/metric_space/hausdorff_distance.lean +++ b/src/topology/metric_space/hausdorff_distance.lean @@ -167,7 +167,7 @@ by simp only [inf_edist, infi_image, hΦ.edist_eq] lemma _root_.is_open.exists_Union_is_closed {U : set α} (hU : is_open U) : ∃ F : ℕ → set α, (∀ n, is_closed (F n)) ∧ (∀ n, F n ⊆ U) ∧ ((⋃ n, F n) = U) ∧ monotone F := begin - obtain ⟨a, a_pos, a_lt_one⟩ : ∃ (a : ℝ≥0∞), 0 < a ∧ a < 1 := exists_between (ennreal.zero_lt_one), + obtain ⟨a, a_pos, a_lt_one⟩ : ∃ (a : ℝ≥0∞), 0 < a ∧ a < 1 := exists_between zero_lt_one, let F := λ (n : ℕ), (λ x, inf_edist x Uᶜ) ⁻¹' (Ici (a^n)), have F_subset : ∀ n, F n ⊆ U, { assume n x hx, @@ -178,7 +178,7 @@ begin show monotone F, { assume m n hmn x hx, simp only [mem_Ici, mem_preimage] at hx ⊢, - apply le_trans (ennreal.pow_le_pow_of_le_one a_lt_one.le hmn) hx }, + apply le_trans (pow_le_pow_of_le_one' a_lt_one.le hmn) hx }, show (⋃ n, F n) = U, { refine subset.antisymm (by simp only [Union_subset_iff, F_subset, forall_const]) (λ x hx, _), have : ¬(x ∈ Uᶜ), by simpa using hx, @@ -1111,7 +1111,7 @@ begin add_le_add hyy'.le $ edist_le_diam_of_mem hy' hx').trans_eq _), -- Now we're done, but `ring` won't do it because we're on `ennreal` :( rw [←add_assoc, ←two_mul, mul_add, - ennreal.mul_div_cancel' ennreal.two_ne_zero ennreal.two_ne_top], + ennreal.mul_div_cancel' two_ne_zero ennreal.two_ne_top], abel, end diff --git a/src/topology/metric_space/metric_separated.lean b/src/topology/metric_space/metric_separated.lean index 8e6bbf6f41d6c..efd1f2d368e7b 100644 --- a/src/topology/metric_space/metric_separated.lean +++ b/src/topology/metric_space/metric_separated.lean @@ -33,7 +33,7 @@ let ⟨r, r0, hr⟩ := h in ⟨r, r0, λ y hy x hx, edist_comm x y ▸ hr x hx y lemma comm : is_metric_separated s t ↔ is_metric_separated t s := ⟨symm, symm⟩ @[simp] lemma empty_left (s : set X) : is_metric_separated ∅ s := -⟨1, ennreal.zero_lt_one.ne', λ x, false.elim⟩ +⟨1, one_ne_zero, λ x, false.elim⟩ @[simp] lemma empty_right (s : set X) : is_metric_separated s ∅ := (empty_left s).symm From 114ff8a4a7935cb7531062200bff375e7b1d6d85 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Mon, 27 Feb 2023 10:42:11 +0000 Subject: [PATCH 0545/1291] feat(data/nat/multiplicity): rename `nat.multiplicity_choose_prime_pow`, add lemmas (#18183) * Rename `nat.multiplicity_choose_prime_pow` to `nat.multiplicity_choose_prime_pow_add_multiplicity`. * Add `part_enat.eq_coe_sub_of_add_eq_coe` and a version of `nat.multiplicity_choose_prime_pow` with just the multiplicity of `p` in `choose (p ^ n) k` in the LHS. * Golf 2 proofs. --- src/data/nat/multiplicity.lean | 11 ++++-- src/data/nat/part_enat.lean | 11 ++++++ src/ring_theory/witt_vector/frobenius.lean | 40 +++++++--------------- 3 files changed, 32 insertions(+), 30 deletions(-) diff --git a/src/data/nat/multiplicity.lean b/src/data/nat/multiplicity.lean index 58f3365ed3fa0..0f1f39d09b6f7 100644 --- a/src/data/nat/multiplicity.lean +++ b/src/data/nat/multiplicity.lean @@ -203,8 +203,8 @@ begin exact dvd_mul_right _ _ end -lemma multiplicity_choose_prime_pow {p n k : ℕ} (hp : p.prime) - (hkn : k ≤ p ^ n) (hk0 : 0 < k) : +lemma multiplicity_choose_prime_pow_add_multiplicity {p n k : ℕ} (hp : p.prime) (hkn : k ≤ p ^ n) + (hk0 : k ≠ 0) : multiplicity p (choose (p ^ n) k) + multiplicity p k = n := le_antisymm (have hdisj : disjoint @@ -214,7 +214,7 @@ le_antisymm {contextual := tt}, begin rw [multiplicity_choose hp hkn (lt_succ_self _), - multiplicity_eq_card_pow_dvd (ne_of_gt hp.one_lt) hk0 + multiplicity_eq_card_pow_dvd (ne_of_gt hp.one_lt) hk0.bot_lt (lt_succ_of_le (log_mono_right hkn)), ← nat.cast_add, part_enat.coe_le_coe, log_pow hp.one_lt, ← card_disjoint_union hdisj, filter_union_right], @@ -224,6 +224,11 @@ le_antisymm (by rw [← hp.multiplicity_pow_self]; exact multiplicity_le_multiplicity_choose_add hp _ _) +lemma multiplicity_choose_prime_pow {p n k : ℕ} (hp : p.prime) (hkn : k ≤ p ^ n) (hk0 : k ≠ 0) : + multiplicity p (choose (p ^ n) k) = + ↑(n - (multiplicity p k).get (finite_nat_iff.2 ⟨hp.ne_one, hk0.bot_lt⟩)) := +part_enat.eq_coe_sub_of_add_eq_coe $ multiplicity_choose_prime_pow_add_multiplicity hp hkn hk0 + end prime lemma multiplicity_two_factorial_lt : ∀ {n : ℕ} (h : n ≠ 0), multiplicity 2 n! < n := diff --git a/src/data/nat/part_enat.lean b/src/data/nat/part_enat.lean index 7ce40ec7e0dea..7f03e53a943cf 100644 --- a/src/data/nat/part_enat.lean +++ b/src/data/nat/part_enat.lean @@ -94,6 +94,8 @@ lemma some_eq_coe (n : ℕ) : some n = n := rfl @[simp] lemma dom_coe (x : ℕ) : (x : part_enat).dom := trivial +instance : can_lift part_enat ℕ coe dom := ⟨λ n hn, ⟨n.get hn, part.some_get _⟩⟩ + instance : has_le part_enat := ⟨λ x y, ∃ h : y.dom → x.dom, ∀ hy : y.dom, x.get (h hy) ≤ y.get hy⟩ instance : has_top part_enat := ⟨none⟩ instance : has_bot part_enat := ⟨0⟩ @@ -326,6 +328,15 @@ instance : canonically_ordered_add_monoid part_enat := ..part_enat.order_bot, ..part_enat.ordered_add_comm_monoid } +lemma eq_coe_sub_of_add_eq_coe {x y : part_enat} {n : ℕ} (h : x + y = n) : + x = ↑(n - y.get (dom_of_le_coe ((le_add_left le_rfl).trans_eq h))) := +begin + lift x to ℕ using dom_of_le_coe ((le_add_right le_rfl).trans_eq h), + lift y to ℕ using dom_of_le_coe ((le_add_left le_rfl).trans_eq h), + rw [← nat.cast_add, coe_inj] at h, + rw [get_coe, coe_inj, eq_tsub_of_add_eq h] +end + protected lemma add_lt_add_right {x y z : part_enat} (h : x < y) (hz : z ≠ ⊤) : x + z < y + z := begin rcases ne_top_iff.mp (ne_top_of_lt h) with ⟨m, rfl⟩, diff --git a/src/ring_theory/witt_vector/frobenius.lean b/src/ring_theory/witt_vector/frobenius.lean index 08f61916e27e0..c2d56ed005ce4 100644 --- a/src/ring_theory/witt_vector/frobenius.lean +++ b/src/ring_theory/witt_vector/frobenius.lean @@ -121,37 +121,23 @@ lemma map_frobenius_poly.key₁ (n j : ℕ) (hj : j < p ^ (n)) : p ^ (n - v p ⟨j + 1, j.succ_pos⟩) ∣ (p ^ n).choose (j + 1) := begin apply multiplicity.pow_dvd_of_le_multiplicity, - have aux : (multiplicity p ((p ^ n).choose (j + 1))).dom, - { rw [← multiplicity.finite_iff_dom, multiplicity.finite_nat_iff], - exact ⟨hp.1.ne_one, nat.choose_pos hj⟩, }, - rw [← part_enat.coe_get aux, part_enat.coe_le_coe, tsub_le_iff_left, - ← part_enat.coe_le_coe, nat.cast_add, pnat_multiplicity, part_enat.coe_get, - part_enat.coe_get, add_comm], - exact (hp.1.multiplicity_choose_prime_pow hj j.succ_pos).ge, + rw [hp.out.multiplicity_choose_prime_pow hj j.succ_ne_zero], + refl, end /-- A key numerical identity needed for the proof of `witt_vector.map_frobenius_poly`. -/ -lemma map_frobenius_poly.key₂ {n i j : ℕ} (hi : i < n) (hj : j < p ^ (n - i)) : - j - (v p ⟨j + 1, j.succ_pos⟩) + n = - i + j + (n - i - v p ⟨j + 1, j.succ_pos⟩) := +lemma map_frobenius_poly.key₂ {n i j : ℕ} (hi : i ≤ n) (hj : j < p ^ (n - i)) : + j - v p ⟨j + 1, j.succ_pos⟩ + n = i + j + (n - i - v p ⟨j + 1, j.succ_pos⟩) := begin generalize h : (v p ⟨j + 1, j.succ_pos⟩) = m, - suffices : m ≤ n - i ∧ m ≤ j, - { rw [tsub_add_eq_add_tsub this.2, add_comm i j, - add_tsub_assoc_of_le (this.1.trans (nat.sub_le n i)), add_assoc, tsub_right_comm, add_comm i, - tsub_add_cancel_of_le (le_tsub_of_add_le_right ((le_tsub_iff_left hi.le).mp this.1))] }, - split, - { rw [← h, ← part_enat.coe_le_coe, pnat_multiplicity, part_enat.coe_get, - ← hp.1.multiplicity_choose_prime_pow hj j.succ_pos], - apply le_add_left, refl }, - { obtain ⟨c, hc⟩ : p ^ m ∣ j + 1, - { rw [← h], exact multiplicity.pow_multiplicity_dvd _, }, - obtain ⟨c, rfl⟩ : ∃ k : ℕ, c = k + 1, - { apply nat.exists_eq_succ_of_ne_zero, rintro rfl, simpa only using hc }, - rw [mul_add, mul_one] at hc, - apply nat.le_of_lt_succ, - calc m < p ^ m : nat.lt_pow_self hp.1.one_lt m - ... ≤ j + 1 : by { rw ← tsub_eq_of_eq_add_rev hc, apply nat.sub_le } } + rsuffices ⟨h₁, h₂⟩ : m ≤ n - i ∧ m ≤ j, + { rw [tsub_add_eq_add_tsub h₂, add_comm i j, + add_tsub_assoc_of_le (h₁.trans (nat.sub_le n i)), add_assoc, tsub_right_comm, add_comm i, + tsub_add_cancel_of_le (le_tsub_of_add_le_right ((le_tsub_iff_left hi).mp h₁))] }, + have hle : p ^ m ≤ j + 1, + from h ▸ nat.le_of_dvd j.succ_pos (multiplicity.pow_multiplicity_dvd _), + exact ⟨(pow_le_pow_iff hp.1.one_lt).1 (hle.trans hj), + nat.le_of_lt_succ ((nat.lt_pow_self hp.1.one_lt m).trans_le hle)⟩ end lemma map_frobenius_poly (n : ℕ) : @@ -200,7 +186,7 @@ begin { have aux : ∀ k : ℕ, (p ^ k : ℚ) ≠ 0, { intro, apply pow_ne_zero, exact_mod_cast hp.1.ne_zero }, simpa [aux, -one_div] with field_simps using this.symm }, - rw [mul_comm _ (p : ℚ), mul_assoc, mul_assoc, ← pow_add, map_frobenius_poly.key₂ p hi hj], + rw [mul_comm _ (p : ℚ), mul_assoc, mul_assoc, ← pow_add, map_frobenius_poly.key₂ p hi.le hj], ring_exp end From f91d3736cbfcff4614caabcf9ae32d4a880f8813 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 27 Feb 2023 10:42:12 +0000 Subject: [PATCH 0546/1291] feat(linear_algebra/bilinear_form/tensor_product): tensor product of bilinear forms (#18211) --- .../bilinear_form/tensor_product.lean | 83 +++++++++++++++++++ 1 file changed, 83 insertions(+) create mode 100644 src/linear_algebra/bilinear_form/tensor_product.lean diff --git a/src/linear_algebra/bilinear_form/tensor_product.lean b/src/linear_algebra/bilinear_form/tensor_product.lean new file mode 100644 index 0000000000000..fb28fadf3d5d4 --- /dev/null +++ b/src/linear_algebra/bilinear_form/tensor_product.lean @@ -0,0 +1,83 @@ +/- +Copyright (c) 2023 Eric Wieser. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Eric Wieser +-/ +import linear_algebra.bilinear_form +import linear_algebra.tensor_product +import linear_algebra.contraction + +/-! +# The bilinear form on a tensor product + +## Main definitions + +* `bilin_form.tensor_distrib (B₁ ⊗ₜ B₂)`: the bilinear form on `M₁ ⊗ M₂` constructed by applying + `B₁` on `M₁` and `B₂` on `M₂`. +* `bilin_form.tensor_distrib_equiv`: `bilin_form.tensor_distrib` as an equivalence on finite free + modules. + +-/ + +universes u v w +variables {ι : Type*} {R : Type*} {M₁ M₂ : Type*} + +open_locale tensor_product + +namespace bilin_form + +section comm_semiring +variables [comm_semiring R] +variables [add_comm_monoid M₁] [add_comm_monoid M₂] +variables [module R M₁] [module R M₂] + +/-- The tensor product of two bilinear forms injects into bilinear forms on tensor products. -/ +def tensor_distrib : bilin_form R M₁ ⊗[R] bilin_form R M₂ →ₗ[R] bilin_form R (M₁ ⊗[R] M₂) := +((tensor_product.tensor_tensor_tensor_comm R _ _ _ _).dual_map + ≪≫ₗ (tensor_product.lift.equiv R _ _ _).symm + ≪≫ₗ linear_map.to_bilin).to_linear_map + ∘ₗ tensor_product.dual_distrib R _ _ + ∘ₗ (tensor_product.congr + (bilin_form.to_lin ≪≫ₗ tensor_product.lift.equiv R _ _ _) + (bilin_form.to_lin ≪≫ₗ tensor_product.lift.equiv R _ _ _)).to_linear_map + +@[simp] lemma tensor_distrib_tmul (B₁ : bilin_form R M₁) (B₂ : bilin_form R M₂) + (m₁ : M₁) (m₂ : M₂) (m₁' : M₁) (m₂' : M₂) : + tensor_distrib (B₁ ⊗ₜ B₂) (m₁ ⊗ₜ m₂) (m₁' ⊗ₜ m₂') = B₁ m₁ m₁' * B₂ m₂ m₂' := +rfl + +/-- The tensor product of two bilinear forms, a shorthand for dot notation. -/ +@[reducible] +protected def tmul (B₁ : bilin_form R M₁) (B₂ : bilin_form R M₂) : bilin_form R (M₁ ⊗[R] M₂) := +tensor_distrib (B₁ ⊗ₜ[R] B₂) + +end comm_semiring + +section comm_ring +variables [comm_ring R] +variables [add_comm_group M₁] [add_comm_group M₂] +variables [module R M₁] [module R M₂] +variables [module.free R M₁] [module.finite R M₁] +variables [module.free R M₂] [module.finite R M₂] +variables [nontrivial R] + +/-- `tensor_distrib` as an equivalence. -/ +noncomputable def tensor_distrib_equiv : + bilin_form R M₁ ⊗[R] bilin_form R M₂ ≃ₗ[R] bilin_form R (M₁ ⊗[R] M₂) := +-- the same `linear_equiv`s as from `tensor_distrib`, but with the inner linear map also as an +-- equiv +tensor_product.congr + (bilin_form.to_lin ≪≫ₗ tensor_product.lift.equiv R _ _ _) + (bilin_form.to_lin ≪≫ₗ tensor_product.lift.equiv R _ _ _) + ≪≫ₗ tensor_product.dual_distrib_equiv R (M₁ ⊗ M₁) (M₂ ⊗ M₂) + ≪≫ₗ (tensor_product.tensor_tensor_tensor_comm R _ _ _ _).dual_map + ≪≫ₗ (tensor_product.lift.equiv R _ _ _).symm + ≪≫ₗ linear_map.to_bilin + +@[simp] +lemma tensor_distrib_equiv_apply (B : bilin_form R M₁ ⊗ bilin_form R M₂) : + tensor_distrib_equiv B = tensor_distrib B := rfl + +end comm_ring + +end bilin_form From aedfb56fb0c32c0146e5dd15f5a1c5fe2088d15a Mon Sep 17 00:00:00 2001 From: JovanGerb Date: Mon, 27 Feb 2023 10:42:13 +0000 Subject: [PATCH 0547/1291] feat(geometry/euclidean/angle/unoriented): sine & cosine angle lemmas (#18222) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit add some lemma's about sine and cosine of unoriented angles. Co-authored-by: Yaël Dillies --- .../euclidean/angle/unoriented/affine.lean | 47 +++++++++++++++++++ .../euclidean/angle/unoriented/basic.lean | 38 ++++++++++++++- 2 files changed, 84 insertions(+), 1 deletion(-) diff --git a/src/geometry/euclidean/angle/unoriented/affine.lean b/src/geometry/euclidean/angle/unoriented/affine.lean index c273b56f61253..9a0025136b94f 100644 --- a/src/geometry/euclidean/angle/unoriented/affine.lean +++ b/src/geometry/euclidean/angle/unoriented/affine.lean @@ -439,4 +439,51 @@ lemma angle_lt_pi_of_not_collinear {p₁ p₂ p₃ : P} (h : ¬collinear ℝ ({p ∠ p₁ p₂ p₃ < π := (angle_le_pi _ _ _).lt_of_ne $ angle_ne_pi_of_not_collinear h +/-- The cosine of the angle between three points is 1 if and only if the angle is 0. -/ +lemma cos_eq_one_iff_angle_eq_zero {p₁ p₂ p₃ : P} : + real.cos (∠ p₁ p₂ p₃) = 1 ↔ ∠ p₁ p₂ p₃ = 0 := +cos_eq_one_iff_angle_eq_zero + +/-- The cosine of the angle between three points is 0 if and only if the angle is π / 2. -/ +lemma cos_eq_zero_iff_angle_eq_pi_div_two {p₁ p₂ p₃ : P} : + real.cos (∠ p₁ p₂ p₃) = 0 ↔ ∠ p₁ p₂ p₃ = π / 2 := +cos_eq_zero_iff_angle_eq_pi_div_two + +/-- The cosine of the angle between three points is -1 if and only if the angle is π. -/ +lemma cos_eq_neg_one_iff_angle_eq_pi {p₁ p₂ p₃ : P} : + real.cos (∠ p₁ p₂ p₃) = -1 ↔ ∠ p₁ p₂ p₃ = π := +cos_eq_neg_one_iff_angle_eq_pi + +/-- The sine of the angle between three points is 0 if and only if the angle is 0 or π. -/ +lemma sin_eq_zero_iff_angle_eq_zero_or_angle_eq_pi {p₁ p₂ p₃ : P} : + real.sin (∠ p₁ p₂ p₃) = 0 ↔ ∠ p₁ p₂ p₃ = 0 ∨ ∠ p₁ p₂ p₃ = π := +sin_eq_zero_iff_angle_eq_zero_or_angle_eq_pi + +/-- The sine of the angle between three points is 1 if and only if the angle is π / 2. -/ +lemma sin_eq_one_iff_angle_eq_pi_div_two {p₁ p₂ p₃ : P} : + real.sin (∠ p₁ p₂ p₃) = 1 ↔ ∠ p₁ p₂ p₃ = π / 2 := +sin_eq_one_iff_angle_eq_pi_div_two + +/-- Three points are collinear if and only if the first or third point equals the second or +the sine of the angle between three points is zero. -/ +lemma collinear_iff_eq_or_eq_or_sin_eq_zero {p₁ p₂ p₃ : P} : + collinear ℝ ({p₁, p₂, p₃} : set P) ↔ p₁ = p₂ ∨ p₃ = p₂ ∨ real.sin (∠ p₁ p₂ p₃) = 0 := +by rw [sin_eq_zero_iff_angle_eq_zero_or_angle_eq_pi, + collinear_iff_eq_or_eq_or_angle_eq_zero_or_angle_eq_pi] + +/-- If three points are not collinear, the sine of the angle between them is positive. -/ +lemma sin_pos_of_not_collinear {p₁ p₂ p₃ : P} (h : ¬collinear ℝ ({p₁, p₂, p₃} : set P)) : + 0 < real.sin (∠ p₁ p₂ p₃) := +real.sin_pos_of_pos_of_lt_pi (angle_pos_of_not_collinear h) (angle_lt_pi_of_not_collinear h) + +/-- If three points are not collinear, the sine of the angle between them is nonzero. -/ +lemma sin_ne_zero_of_not_collinear {p₁ p₂ p₃ : P} (h : ¬collinear ℝ ({p₁, p₂, p₃} : set P)) : + real.sin (∠ p₁ p₂ p₃) ≠ 0 := +ne_of_gt (sin_pos_of_not_collinear h) + +/-- If the sine of the angle between three points is 0, they are collinear. -/ +lemma collinear_of_sin_eq_zero {p₁ p₂ p₃ : P} (h : real.sin (∠ p₁ p₂ p₃) = 0) : + collinear ℝ ({p₁, p₂, p₃} : set P) := +imp_of_not_imp_not _ _ sin_ne_zero_of_not_collinear h + end euclidean_geometry diff --git a/src/geometry/euclidean/angle/unoriented/basic.lean b/src/geometry/euclidean/angle/unoriented/basic.lean index 6afbd85c4ec5c..fd4c805d90234 100644 --- a/src/geometry/euclidean/angle/unoriented/basic.lean +++ b/src/geometry/euclidean/angle/unoriented/basic.lean @@ -21,13 +21,14 @@ assert_not_exists has_fderiv_at assert_not_exists conformal_at noncomputable theory +open real set open_locale big_operators open_locale real open_locale real_inner_product_space namespace inner_product_geometry -variables {V : Type*} [inner_product_space ℝ V] +variables {V : Type*} [inner_product_space ℝ V] {x y : V} /-- The undirected angle between two vectors. If either vector is 0, this is π/2. See `orientation.oangle` for the corresponding oriented angle @@ -317,4 +318,39 @@ begin split; intro h; linarith, end +/-- The cosine of the angle between two vectors is 1 if and only if the angle is 0. -/ +lemma cos_eq_one_iff_angle_eq_zero : cos (angle x y) = 1 ↔ angle x y = 0 := +begin + rw ← cos_zero, + exact inj_on_cos.eq_iff ⟨angle_nonneg x y, angle_le_pi x y⟩ (left_mem_Icc.2 pi_pos.le), +end + +/-- The cosine of the angle between two vectors is 0 if and only if the angle is π / 2. -/ +lemma cos_eq_zero_iff_angle_eq_pi_div_two : cos (angle x y) = 0 ↔ angle x y = π / 2 := +begin + rw ← cos_pi_div_two, + apply inj_on_cos.eq_iff ⟨angle_nonneg x y, angle_le_pi x y⟩, + split; linarith [pi_pos], +end + +/-- The cosine of the angle between two vectors is -1 if and only if the angle is π. -/ +lemma cos_eq_neg_one_iff_angle_eq_pi : cos (angle x y) = -1 ↔ angle x y = π := +begin + rw ← cos_pi, + exact inj_on_cos.eq_iff ⟨angle_nonneg x y, angle_le_pi x y⟩ (right_mem_Icc.2 pi_pos.le), +end + +/-- The sine of the angle between two vectors is 0 if and only if the angle is 0 or π. -/ +lemma sin_eq_zero_iff_angle_eq_zero_or_angle_eq_pi : + sin (angle x y) = 0 ↔ angle x y = 0 ∨ angle x y = π := +by rw [sin_eq_zero_iff_cos_eq, cos_eq_one_iff_angle_eq_zero, cos_eq_neg_one_iff_angle_eq_pi] + +/-- The sine of the angle between two vectors is 1 if and only if the angle is π / 2. -/ +lemma sin_eq_one_iff_angle_eq_pi_div_two : sin (angle x y) = 1 ↔ angle x y = π / 2 := +begin + refine ⟨λ h, _, λ h, by rw [h, sin_pi_div_two]⟩, + rw [←cos_eq_zero_iff_angle_eq_pi_div_two, ←abs_eq_zero, abs_cos_eq_sqrt_one_sub_sin_sq, h], + simp, +end + end inner_product_geometry From c3fc15b26b3ff8958ec3e5711177a9ae3d5c45b7 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 27 Feb 2023 10:42:14 +0000 Subject: [PATCH 0548/1291] fix(data/finmap): `finmap.all` should not always be `ff` (#18432) Also removes some unnecesary back and forth between bool and Prop. --- src/data/finmap.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/data/finmap.lean b/src/data/finmap.lean index 2fb552f8cfdf3..0f49f8d4311e9 100644 --- a/src/data/finmap.lean +++ b/src/data/finmap.lean @@ -252,11 +252,11 @@ m.entries.foldl (λ d s, f d s.1 s.2) (λ d s t, H _ _ _ _ _) d /-- `any f s` returns `tt` iff there exists a value `v` in `s` such that `f v = tt`. -/ def any (f : Π x, β x → bool) (s : finmap β) : bool := -s.foldl (λ x y z, x ∨ f y z) (by { intros, simp [or.right_comm] }) ff +s.foldl (λ x y z, x || f y z) (by { intros, simp_rw [bool.bor_assoc, bool.bor_comm] }) ff /-- `all f s` returns `tt` iff `f v = tt` for all values `v` in `s`. -/ def all (f : Π x, β x → bool) (s : finmap β) : bool := -s.foldl (λ x y z, x ∧ f y z) (by { intros, simp [and.right_comm] }) ff +s.foldl (λ x y z, x && f y z) (by { intros, simp_rw [bool.band_assoc, bool.band_comm] }) tt /-! ### erase -/ From 824f9ae93a4f5174d2ea948e2d75843dd83447bb Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Mon, 27 Feb 2023 10:42:15 +0000 Subject: [PATCH 0549/1291] feat(algebra/order/ring/canonical): add `canonically_ordered_comm_semiring.to_ordered_comm_monoid` (#18504) Also merge `finset.prod_le_prod'` with `finset.prod_le_prod''`. Mathlib 4 version is leanprover-community/mathlib4#2510 --- src/algebra/big_operators/order.lean | 28 ++++++++------------------- src/algebra/order/ring/canonical.lean | 5 +++++ src/number_theory/bertrand.lean | 2 +- 3 files changed, 14 insertions(+), 21 deletions(-) diff --git a/src/algebra/big_operators/order.lean b/src/algebra/big_operators/order.lean index 25dd5b41c238d..2573be61bcdb6 100644 --- a/src/algebra/big_operators/order.lean +++ b/src/algebra/big_operators/order.lean @@ -110,7 +110,7 @@ variables {f g : ι → N} {s t : finset ι} equal to the corresponding factor `g i` of another finite product, then `∏ i in s, f i ≤ ∏ i in s, g i`. -/ @[to_additive sum_le_sum] -lemma prod_le_prod'' (h : ∀ i ∈ s, f i ≤ g i) : ∏ i in s, f i ≤ ∏ i in s, g i := +lemma prod_le_prod' (h : ∀ i ∈ s, f i ≤ g i) : ∏ i in s, f i ≤ ∏ i in s, g i := multiset.prod_map_le_prod_map f g h /-- In an ordered additive commutative monoid, if each summand `f i` of one finite sum is less than @@ -119,14 +119,14 @@ or equal to the corresponding summand `g i` of another finite sum, then add_decl_doc sum_le_sum @[to_additive sum_nonneg] lemma one_le_prod' (h : ∀i ∈ s, 1 ≤ f i) : 1 ≤ (∏ i in s, f i) := -le_trans (by rw prod_const_one) (prod_le_prod'' h) +le_trans (by rw prod_const_one) (prod_le_prod' h) @[to_additive finset.sum_nonneg'] lemma one_le_prod'' (h : ∀ (i : ι), 1 ≤ f i) : 1 ≤ ∏ (i : ι) in s, f i := finset.one_le_prod' (λ i hi, h i) @[to_additive sum_nonpos] lemma prod_le_one' (h : ∀i ∈ s, f i ≤ 1) : (∏ i in s, f i) ≤ 1 := -(prod_le_prod'' h).trans_eq (by rw prod_const_one) +(prod_le_prod' h).trans_eq (by rw prod_const_one) @[to_additive sum_le_sum_of_subset_of_nonneg] lemma prod_le_prod_of_subset_of_one_le' (h : s ⊆ t) (hf : ∀ i ∈ t, i ∉ s → 1 ≤ f i) : @@ -367,7 +367,7 @@ begin classical, rcases Hlt with ⟨i, hi, hlt⟩, rw [← insert_erase hi, prod_insert (not_mem_erase _ _), prod_insert (not_mem_erase _ _)], - exact mul_lt_mul_of_lt_of_le hlt (prod_le_prod'' $ λ j hj, Hle j $ mem_of_mem_erase hj) + exact mul_lt_mul_of_lt_of_le hlt (prod_le_prod' $ λ j hj, Hle j $ mem_of_mem_erase hj) end @[to_additive sum_lt_sum_of_nonempty] @@ -432,7 +432,7 @@ begin refine finset.induction_on s (λ _, ⟨λ _ _, false.elim, λ _, rfl⟩) (λ a s ha ih H, _), specialize ih (λ i, H i ∘ finset.mem_insert_of_mem), rw [finset.prod_insert ha, finset.prod_insert ha, finset.forall_mem_insert, ←ih], - exact mul_eq_mul_iff_eq_and_eq (H a (s.mem_insert_self a)) (finset.prod_le_prod'' + exact mul_eq_mul_iff_eq_and_eq (H a (s.mem_insert_self a)) (finset.prod_le_prod' (λ i, H i ∘ finset.mem_insert_of_mem)), end @@ -447,7 +447,7 @@ theorem exists_lt_of_prod_lt' (Hlt : ∏ i in s, f i < ∏ i in s, g i) : ∃ i ∈ s, f i < g i := begin contrapose! Hlt with Hle, - exact prod_le_prod'' Hle + exact prod_le_prod' Hle end @[to_additive exists_le_of_sum_le] @@ -482,7 +482,7 @@ lemma prod_nonneg (h0 : ∀ i ∈ s, 0 ≤ f i) : 0 ≤ ∏ i in s, f i := prod_induction f (λ i, 0 ≤ i) (λ _ _ ha hb, mul_nonneg ha hb) zero_le_one h0 /-- If all `f i`, `i ∈ s`, are nonnegative and each `f i` is less than or equal to `g i`, then the -product of `f i` is less than or equal to the product of `g i`. See also `finset.prod_le_prod''` for +product of `f i` is less than or equal to the product of `g i`. See also `finset.prod_le_prod'` for the case of an ordered commutative multiplicative monoid. -/ lemma prod_le_prod (h0 : ∀ i ∈ s, 0 ≤ f i) (h1 : ∀ i ∈ s, f i ≤ g i) : ∏ i in s, f i ≤ ∏ i in s, g i := @@ -551,18 +551,6 @@ lemma _root_.canonically_ordered_comm_semiring.prod_pos [nontrivial R] : 0 < ∏ i in s, f i ↔ (∀ i ∈ s, (0 : R) < f i) := canonically_ordered_comm_semiring.multiset_prod_pos.trans $ by simp -lemma prod_le_prod' (h : ∀ i ∈ s, f i ≤ g i) : - ∏ i in s, f i ≤ ∏ i in s, g i := -begin - classical, - induction s using finset.induction with a s has ih h, - { simp }, - { rw [finset.prod_insert has, finset.prod_insert has], - apply mul_le_mul', - { exact h _ (finset.mem_insert_self a s) }, - { exact ih (λ i hi, h _ (finset.mem_insert_of_mem hi)) } } -end - /-- If `g, h ≤ f` and `g i + h i ≤ f i`, then the product of `f` over `s` is at least the sum of the products of `g` and `h`. This is the version for `canonically_ordered_comm_semiring`. -/ @@ -587,7 +575,7 @@ variables [fintype ι] @[to_additive sum_mono, mono] lemma prod_mono' [ordered_comm_monoid M] : monotone (λ f : ι → M, ∏ i, f i) := -λ f g hfg, finset.prod_le_prod'' $ λ x _, hfg x +λ f g hfg, finset.prod_le_prod' $ λ x _, hfg x attribute [mono] sum_mono diff --git a/src/algebra/order/ring/canonical.lean b/src/algebra/order/ring/canonical.lean index 1b4a386f746cd..c9a037c59efd2 100644 --- a/src/algebra/order/ring/canonical.lean +++ b/src/algebra/order/ring/canonical.lean @@ -92,6 +92,11 @@ begin apply self_le_add_right end +@[priority 100] -- see Note [lower instance priority] +instance to_ordered_comm_monoid : ordered_comm_monoid α := +{ mul_le_mul_left := λ _ _, mul_le_mul_left', + .. ‹canonically_ordered_comm_semiring α› } + @[priority 100] -- see Note [lower instance priority] instance to_ordered_comm_semiring : ordered_comm_semiring α := { zero_le_one := zero_le _, diff --git a/src/number_theory/bertrand.lean b/src/number_theory/bertrand.lean index a3cc2fd71c6d6..b9dbb2af60d55 100644 --- a/src/number_theory/bertrand.lean +++ b/src/number_theory/bertrand.lean @@ -156,7 +156,7 @@ begin rw [central_binom_factorization_small n n_big no_prime, ← this, ← finset.prod_filter_mul_prod_filter_not S (≤ sqrt (2 * n))], apply mul_le_mul', - { refine (finset.prod_le_prod'' (λ p hp, (_ : f p ≤ 2 * n))).trans _, + { refine (finset.prod_le_prod' (λ p hp, (_ : f p ≤ 2 * n))).trans _, { exact pow_factorization_choose_le (mul_pos two_pos n_pos) }, have : (finset.Icc 1 (sqrt (2 * n))).card = sqrt (2 * n), { rw [card_Icc, nat.add_sub_cancel] }, From 35c1956cb727c474aa8863c13ca3abca4c2f885e Mon Sep 17 00:00:00 2001 From: Thomas Browning Date: Mon, 27 Feb 2023 14:28:53 +0000 Subject: [PATCH 0550/1291] refactor(number_theory/pell): Small golf using `int.eq_of_mul_eq_one` (#18505) This is a small golf using `int.eq_of_mul_eq_one`. --- src/number_theory/pell.lean | 5 ++--- 1 file changed, 2 insertions(+), 3 deletions(-) diff --git a/src/number_theory/pell.lean b/src/number_theory/pell.lean index 6d4cccaab4c46..6cd7fc96a9a1d 100644 --- a/src/number_theory/pell.lean +++ b/src/number_theory/pell.lean @@ -118,9 +118,8 @@ theorem exists_iff_not_is_square {d : ℤ} (h₀ : 0 < d) : begin refine ⟨_, exists_of_not_is_square h₀⟩, rintros ⟨x, y, hxy, hy⟩ ⟨a, rfl⟩, - rw [← sq, ← mul_pow, sq_sub_sq, int.mul_eq_one_iff_eq_one_or_neg_one] at hxy, - replace hxy := hxy.elim (λ h, h.1.trans h.2.symm) (λ h, h.1.trans h.2.symm), - simpa [mul_self_pos.mp h₀, sub_eq_add_neg, eq_neg_self_iff] using hxy, + rw [← sq, ← mul_pow, sq_sub_sq] at hxy, + simpa [mul_self_pos.mp h₀, sub_eq_add_neg, eq_neg_self_iff] using int.eq_of_mul_eq_one hxy, end end existence From 9da1b3534b65d9661eb8f42443598a92bbb49211 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 27 Feb 2023 16:30:47 +0000 Subject: [PATCH 0551/1291] feat(data/list/basic): missing lemmas about `list.enum` (#18489) The primed versions of these lemmas are more convenient when inducting over the list argument. Forward-ported at https://github.com/leanprover-community/mathlib4/pull/2469. --- src/data/list/basic.lean | 20 ++++++++++++++++++++ 1 file changed, 20 insertions(+) diff --git a/src/data/list/basic.lean b/src/data/list/basic.lean index 242db1cf4034d..939c431e6d038 100644 --- a/src/data/list/basic.lean +++ b/src/data/list/basic.lean @@ -3704,6 +3704,26 @@ lemma map_fst_add_enum_eq_enum_from (l : list α) (n : ℕ) : map (prod.map (+ n) id) (enum l) = enum_from n l := map_fst_add_enum_from_eq_enum_from l _ _ +lemma enum_from_cons' (n : ℕ) (x : α) (xs : list α) : + enum_from n (x :: xs) = (n, x) :: (enum_from n xs).map (prod.map nat.succ id) := +by rw [enum_from_cons, add_comm, ←map_fst_add_enum_from_eq_enum_from] + +lemma enum_cons' (x : α) (xs : list α) : + enum (x :: xs) = (0, x) :: (enum xs).map (prod.map nat.succ id) := +enum_from_cons' _ _ _ + +lemma enum_from_map (n : ℕ) (l : list α) (f : α → β) : + enum_from n (l.map f) = (enum_from n l).map (prod.map id f) := +begin + induction l with hd tl IH, + { refl }, + { rw [map_cons, enum_from_cons', enum_from_cons', map_cons, map_map, IH, map_map], + refl, }, +end + +lemma enum_map (l : list α) (f : α → β) : (l.map f).enum = l.enum.map (prod.map id f) := +enum_from_map _ _ _ + lemma nth_le_enum_from (l : list α) (n i : ℕ) (hi' : i < (l.enum_from n).length) (hi : i < l.length := by simpa [length_enum_from] using hi') : From 23a3e7d9c5bc567d2c91ea8d92ff8da3d6621094 Mon Sep 17 00:00:00 2001 From: Floris van Doorn Date: Mon, 27 Feb 2023 21:24:31 +0000 Subject: [PATCH 0552/1291] fix(geometry/manifold/vector_bundle): fix timeout (#18507) * Also uniformize the spelling of fiberwise/fibrewise * [Zulip thread](https://leanprover.zulipchat.com/#narrow/stream/113488-general/topic/Timeout.20in.20geometry.2Emanifold.2Evector_bundle) --- src/geometry/manifold/vector_bundle.lean | 86 ++++++++++++++++++------ 1 file changed, 65 insertions(+), 21 deletions(-) diff --git a/src/geometry/manifold/vector_bundle.lean b/src/geometry/manifold/vector_bundle.lean index 6398718dc840f..edc4850291457 100644 --- a/src/geometry/manifold/vector_bundle.lean +++ b/src/geometry/manifold/vector_bundle.lean @@ -9,7 +9,7 @@ import geometry.manifold.cont_mdiff This file will eventually contain the definition of a smooth vector bundle. For now, it contains preliminaries regarding an associated `structure_groupoid`, the groupoid of -`smooth_fibrewise_linear` functions. When a (topological) vector bundle is smooth, then the +`smooth_fiberwise_linear` functions. When a (topological) vector bundle is smooth, then the composition of charts associated to the vector bundle belong to this groupoid. -/ @@ -18,7 +18,7 @@ noncomputable theory open set topological_space open_locale manifold topology -/-! ### The groupoid of smooth, fibrewise-linear maps -/ +/-! ### The groupoid of smooth, fiberwise-linear maps -/ variables {𝕜 B F : Type*} [topological_space B] variables [nontrivially_normed_field 𝕜] [normed_add_comm_group F] [normed_space 𝕜 F] @@ -28,7 +28,7 @@ namespace fiberwise_linear variables {φ φ' : B → F ≃L[𝕜] F} {U U' : set B} /-- For `B` a topological space and `F` a `𝕜`-normed space, a map from `U : set B` to `F ≃L[𝕜] F` -determines a local homeomorphism from `B × F` to itself by its action fibrewise. -/ +determines a local homeomorphism from `B × F` to itself by its action fiberwise. -/ def local_homeomorph (φ : B → F ≃L[𝕜] F) (hU : is_open U) (hφ : continuous_on (λ x, φ x : B → F →L[𝕜] F) U) (h2φ : continuous_on (λ x, (φ x).symm : B → F →L[𝕜] F) U) : @@ -65,7 +65,7 @@ lemma trans_local_homeomorph_apply (h2φ' : continuous_on (λ x, (φ' x).symm : B → F →L[𝕜] F) U') (b : B) (v : F) : (fiberwise_linear.local_homeomorph φ hU hφ h2φ ≫ₕ - fiberwise_linear.local_homeomorph φ' hU' hφ' h2φ') ⟨b, v⟩ = ⟨b, φ' b (φ b v)⟩ := + fiberwise_linear.local_homeomorph φ' hU' hφ' h2φ') ⟨b, v⟩ = ⟨b, φ' b (φ b v)⟩ := rfl /-- Compute the source of the composition of two local homeomorphisms induced by fiberwise linear @@ -78,8 +78,8 @@ lemma source_trans_local_homeomorph (hφ' : continuous_on (λ x, φ' x : B → F →L[𝕜] F) U') (h2φ' : continuous_on (λ x, (φ' x).symm : B → F →L[𝕜] F) U') : (fiberwise_linear.local_homeomorph φ hU hφ h2φ ≫ₕ - fiberwise_linear.local_homeomorph φ' hU' hφ' h2φ').source = (U ∩ U') ×ˢ univ := -by { dsimp [fiberwise_linear.local_homeomorph], mfld_set_tac } + fiberwise_linear.local_homeomorph φ' hU' hφ' h2φ').source = (U ∩ U') ×ˢ univ := +by { dsimp only [fiberwise_linear.local_homeomorph], mfld_set_tac } /-- Compute the target of the composition of two local homeomorphisms induced by fiberwise linear equivalences. -/ @@ -91,8 +91,8 @@ lemma target_trans_local_homeomorph (hφ' : continuous_on (λ x, φ' x : B → F →L[𝕜] F) U') (h2φ' : continuous_on (λ x, (φ' x).symm : B → F →L[𝕜] F) U') : (fiberwise_linear.local_homeomorph φ hU hφ h2φ ≫ₕ - fiberwise_linear.local_homeomorph φ' hU' hφ' h2φ').target = (U ∩ U') ×ˢ univ := -by { dsimp [fiberwise_linear.local_homeomorph], mfld_set_tac } + fiberwise_linear.local_homeomorph φ' hU' hφ' h2φ').target = (U ∩ U') ×ˢ univ := +by { dsimp only [fiberwise_linear.local_homeomorph], mfld_set_tac } end fiberwise_linear @@ -100,19 +100,19 @@ variables {EB : Type*} [normed_add_comm_group EB] [normed_space 𝕜 EB] {HB : Type*} [topological_space HB] [charted_space HB B] {IB : model_with_corners 𝕜 EB HB} /-- Let `e` be a local homeomorphism of `B × F`. Suppose that at every point `p` in the source of -`e`, there is some neighbourhood `s` of `p` on which `e` is equal to a bi-smooth fibrewise linear +`e`, there is some neighbourhood `s` of `p` on which `e` is equal to a bi-smooth fiberwise linear local homeomorphism. Then the source of `e` is of the form `U ×ˢ univ`, for some set `U` in `B`, and, at any point `x` in `U`, admits a neighbourhood `u` of `x` such that `e` is equal on `u ×ˢ univ` to some bi-smooth -fibrewise linear local homeomorphism. -/ -lemma smooth_fibrewise_linear.locality_aux₁ (e : local_homeomorph (B × F) (B × F)) +fiberwise linear local homeomorphism. -/ +lemma smooth_fiberwise_linear.locality_aux₁ (e : local_homeomorph (B × F) (B × F)) (h : ∀ p ∈ e.source, ∃ s : set (B × F), is_open s ∧ p ∈ s ∧ ∃ (φ : B → (F ≃L[𝕜] F)) (u : set B) (hu : is_open u) (hφ : smooth_on IB 𝓘(𝕜, F →L[𝕜] F) (λ x, (φ x : F →L[𝕜] F)) u) (h2φ : smooth_on IB 𝓘(𝕜, F →L[𝕜] F) (λ x, ((φ x).symm : F →L[𝕜] F)) u), (e.restr s).eq_on_source - (fiberwise_linear.local_homeomorph φ hu hφ.continuous_on h2φ.continuous_on)) : + (fiberwise_linear.local_homeomorph φ hu hφ.continuous_on h2φ.continuous_on)) : ∃ (U : set B) (hU : e.source = U ×ˢ univ), ∀ x ∈ U, ∃ (φ : B → (F ≃L[𝕜] F)) (u : set B) (hu : is_open u) (huU : u ⊆ U) (hux : x ∈ u) (hφ : smooth_on IB 𝓘(𝕜, F →L[𝕜] F) (λ x, (φ x : F →L[𝕜] F)) u) @@ -121,7 +121,51 @@ lemma smooth_fibrewise_linear.locality_aux₁ (e : local_homeomorph (B × F) (B (fiberwise_linear.local_homeomorph φ hu hφ.continuous_on h2φ.continuous_on) := begin rw set_coe.forall' at h, - choose s hs hsp φ u hu hφ h2φ heφ using h, + -- choose s hs hsp φ u hu hφ h2φ heφ using h, + -- the following ~40 lines should be `choose s hs hsp φ u hu hφ h2φ heφ using h,` + -- `choose` produces a proof term that takes a long time to type-check by the kernel (it seems) + let s := λ x, classical.some (h x), + have hs : ∀ x, is_open (s x) := λ x, (classical.some_spec (h x)).1, + have hsp : ∀ x : e.source, x.1 ∈ s x := λ x, (classical.some_spec (h x)).2.1, + have h : ∀ x : e.source, ∃ (φ : B → (F ≃L[𝕜] F)) (u : set B) (hu : is_open u) + (hφ : smooth_on IB 𝓘(𝕜, F →L[𝕜] F) (λ x, (φ x : F →L[𝕜] F)) u) + (h2φ : smooth_on IB 𝓘(𝕜, F →L[𝕜] F) (λ x, ((φ x).symm : F →L[𝕜] F)) u), + (e.restr (s x)).eq_on_source + (fiberwise_linear.local_homeomorph φ hu hφ.continuous_on h2φ.continuous_on) := + λ x, (classical.some_spec (h x)).2.2, + let φ : e.source → B → (F ≃L[𝕜] F) := λ x, classical.some (h x), + have h : ∀ z : e.source, ∃ (u : set B) (hu : is_open u) + (hφ : smooth_on IB 𝓘(𝕜, F →L[𝕜] F) (λ x, (φ z x : F →L[𝕜] F)) u) + (h2φ : smooth_on IB 𝓘(𝕜, F →L[𝕜] F) (λ x, ((φ z x).symm : F →L[𝕜] F)) u), + (e.restr (s z)).eq_on_source + (fiberwise_linear.local_homeomorph (φ z) hu hφ.continuous_on h2φ.continuous_on) := + λ x, classical.some_spec (h x), + let u : e.source → set B := λ x, classical.some (h x), + have h : ∀ z : e.source, ∃ (hu : is_open (u z)) + (hφ : smooth_on IB 𝓘(𝕜, F →L[𝕜] F) (λ x, (φ z x : F →L[𝕜] F)) (u z)) + (h2φ : smooth_on IB 𝓘(𝕜, F →L[𝕜] F) (λ x, ((φ z x).symm : F →L[𝕜] F)) (u z)), + (e.restr (s z)).eq_on_source + (fiberwise_linear.local_homeomorph (φ z) hu hφ.continuous_on h2φ.continuous_on) := + λ x, classical.some_spec (h x), + have hu : ∀ x, is_open (u x) := λ x, classical.some (h x), + have h : ∀ z : e.source, ∃ (hφ : smooth_on IB 𝓘(𝕜, F →L[𝕜] F) (λ x, (φ z x : F →L[𝕜] F)) (u z)) + (h2φ : smooth_on IB 𝓘(𝕜, F →L[𝕜] F) (λ x, ((φ z x).symm : F →L[𝕜] F)) (u z)), + (e.restr (s z)).eq_on_source + (fiberwise_linear.local_homeomorph (φ z) (hu z) hφ.continuous_on h2φ.continuous_on) := + λ x, classical.some_spec (h x), + have hφ : ∀ z, smooth_on IB 𝓘(𝕜, F →L[𝕜] F) (λ x, (φ z x : F →L[𝕜] F)) (u z) := + λ x, classical.some (h x), + have h : ∀ z : e.source, + ∃ (h2φ : smooth_on IB 𝓘(𝕜, F →L[𝕜] F) (λ x, ((φ z x).symm : F →L[𝕜] F)) (u z)), + (e.restr (s z)).eq_on_source + (fiberwise_linear.local_homeomorph (φ z) (hu z) (hφ z).continuous_on h2φ.continuous_on) := + λ x, classical.some_spec (h x), + have h2φ : ∀ z, smooth_on IB 𝓘(𝕜, F →L[𝕜] F) (λ x, ((φ z x).symm : F →L[𝕜] F)) (u z) := + λ x, classical.some (h x), + have heφ : ∀ z, (e.restr (s z)).eq_on_source + (fiberwise_linear.local_homeomorph (φ z) (hu z) (hφ z).continuous_on (h2φ z).continuous_on) := + λ x, classical.some_spec (h x), + clear_value s φ u, clear h h h h h h, have hesu : ∀ p : e.source, e.source ∩ s p = u p ×ˢ univ, { intros p, rw ← e.restr_source' (s _) (hs _), @@ -150,16 +194,16 @@ end /-- Let `e` be a local homeomorphism of `B × F` whose source is `U ×ˢ univ`, for some set `U` in `B`, and which, at any point `x` in `U`, admits a neighbourhood `u` of `x` such that `e` is equal on -`u ×ˢ univ` to some bi-smooth fibrewise linear local homeomorphism. Then `e` itself is equal to -some bi-smooth fibrewise linear local homeomorphism. +`u ×ˢ univ` to some bi-smooth fiberwise linear local homeomorphism. Then `e` itself is equal to +some bi-smooth fiberwise linear local homeomorphism. This is the key mathematical point of the `locality` condition in the construction of the -`structure_groupoid` of bi-smooth fibrewise linear local homeomorphisms. The proof is by gluing -together the various bi-smooth fibrewise linear local homeomorphism which exist locally. +`structure_groupoid` of bi-smooth fiberwise linear local homeomorphisms. The proof is by gluing +together the various bi-smooth fiberwise linear local homeomorphism which exist locally. The `U` in the conclusion is the same `U` as in the hypothesis. We state it like this, because this is exactly what we need for `smooth_fiberwise_linear`. -/ -lemma smooth_fibrewise_linear.locality_aux₂ (e : local_homeomorph (B × F) (B × F)) +lemma smooth_fiberwise_linear.locality_aux₂ (e : local_homeomorph (B × F) (B × F)) (U : set B) (hU : e.source = U ×ˢ univ) (h : ∀ x ∈ U, ∃ (φ : B → (F ≃L[𝕜] F)) (u : set B) (hu : is_open u) (hUu : u ⊆ U) (hux : x ∈ u) (hφ : smooth_on IB 𝓘(𝕜, F →L[𝕜] F) (λ x, (φ x : F →L[𝕜] F)) u) @@ -226,7 +270,7 @@ end variables (F B IB) /-- For `B` a manifold and `F` a normed space, the groupoid on `B × F` consisting of local -homeomorphisms which are bi-smooth and fibrewise linear, and induce the identity on `B`. +homeomorphisms which are bi-smooth and fiberwise linear, and induce the identity on `B`. When a (topological) vector bundle is smooth, then the composition of charts associated to the vector bundle belong to this groupoid. -/ def smooth_fiberwise_linear : structure_groupoid (B × F) := @@ -267,8 +311,8 @@ def smooth_fiberwise_linear : structure_groupoid (B × F) := locality' := begin -- the hard work has been extracted to `locality_aux₁` and `locality_aux₂` simp_rw [mem_Union], intros e he, - obtain ⟨U, hU, h⟩ := smooth_fibrewise_linear.locality_aux₁ e he, - exact smooth_fibrewise_linear.locality_aux₂ e U hU h, + obtain ⟨U, hU, h⟩ := smooth_fiberwise_linear.locality_aux₁ e he, + exact smooth_fiberwise_linear.locality_aux₂ e U hU h, end, eq_on_source' := begin simp_rw [mem_Union], From e1a7bdeb4fd826b7e71d130d34988f0a2d26a177 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Mon, 27 Feb 2023 22:32:09 +0000 Subject: [PATCH 0553/1291] refactor(topology/*): add `uniform_space.of_fun`, use it (#18495) * Fix `simps` config for `absolute_value`. * Define `uniform_space.of_fun` and use it for `absolute_value.uniform_space`, `pseudo_emetric_space`, and `pseudo_metric_space`. * Add `filter.tendsto_infi_infi` and `filter.tendsto_supr_supr`. * Rename `pseudo_metric_space.of_metrizable` and `metric_space.of_metrizable` to `*.of_dist_topology`. * Add `metric.to_uniform_space_eq` and `metric.uniformity_basis_dist_rat`. * Migrate `topology.uniform_space.absolute_value` to bundled `absolute_value`. --- src/algebra/order/absolute_value.lean | 7 +- src/order/filter/basic.lean | 8 ++ src/topology/metric_space/basic.lean | 78 ++++++------------- src/topology/metric_space/emetric_space.lean | 67 ++++------------ src/topology/metric_space/gluing.lean | 2 +- src/topology/metric_space/pi_nat.lean | 2 +- .../uniform_space/absolute_value.lean | 51 ++++-------- src/topology/uniform_space/basic.lean | 26 +++++++ src/topology/uniform_space/compare_reals.lean | 23 ++---- 9 files changed, 96 insertions(+), 168 deletions(-) diff --git a/src/algebra/order/absolute_value.lean b/src/algebra/order/absolute_value.lean index 906537205350e..fef733c6dbf2e 100644 --- a/src/algebra/order/absolute_value.lean +++ b/src/algebra/order/absolute_value.lean @@ -39,8 +39,6 @@ namespace absolute_value attribute [nolint doc_blame] absolute_value.to_mul_hom -initialize_simps_projections absolute_value (to_mul_hom_to_fun → apply) - section ordered_semiring section semiring @@ -68,6 +66,11 @@ instance subadditive_hom_class : subadditive_hom_class (absolute_value R S) R S @[ext] lemma ext ⦃f g : absolute_value R S⦄ : (∀ x, f x = g x) → f = g := fun_like.ext _ _ +/-- See Note [custom simps projection]. -/ +def simps.apply (f : absolute_value R S) : R → S := f + +initialize_simps_projections absolute_value (to_mul_hom_to_fun → apply) + /-- Helper instance for when there's too many metavariables to apply `fun_like.has_coe_to_fun` directly. -/ instance : has_coe_to_fun (absolute_value R S) (λ f, R → S) := fun_like.has_coe_to_fun diff --git a/src/order/filter/basic.lean b/src/order/filter/basic.lean index 64c0d6fbd66e5..6748130b7d2ef 100644 --- a/src/order/filter/basic.lean +++ b/src/order/filter/basic.lean @@ -2531,6 +2531,10 @@ lemma tendsto_infi' {f : α → β} {x : ι → filter α} {y : filter β} (i : tendsto f (⨅ i, x i) y := hi.mono_left $ infi_le _ _ +theorem tendsto_infi_infi {f : α → β} {x : ι → filter α} {y : ι → filter β} + (h : ∀ i, tendsto f (x i) (y i)) : tendsto f (infi x) (infi y) := +tendsto_infi.2 $ λ i, tendsto_infi' i (h i) + @[simp] lemma tendsto_sup {f : α → β} {x₁ x₂ : filter α} {y : filter β} : tendsto f (x₁ ⊔ x₂) y ↔ tendsto f x₁ y ∧ tendsto f x₂ y := by simp only [tendsto, map_sup, sup_le_iff] @@ -2543,6 +2547,10 @@ lemma tendsto.sup {f : α → β} {x₁ x₂ : filter α} {y : filter β} : tendsto f (⨆ i, x i) y ↔ ∀ i, tendsto f (x i) y := by simp only [tendsto, map_supr, supr_le_iff] +theorem tendsto_supr_supr {f : α → β} {x : ι → filter α} {y : ι → filter β} + (h : ∀ i, tendsto f (x i) (y i)) : tendsto f (supr x) (supr y) := +tendsto_supr.2 $ λ i, (h i).mono_right $ le_supr _ _ + @[simp] lemma tendsto_principal {f : α → β} {l : filter α} {s : set β} : tendsto f l (𝓟 s) ↔ ∀ᶠ a in l, f a ∈ s := by simp only [tendsto, le_principal_iff, mem_map', filter.eventually] diff --git a/src/topology/metric_space/basic.lean b/src/topology/metric_space/basic.lean index c5566c9b5ee0d..e49dcb84c0c74 100644 --- a/src/topology/metric_space/basic.lean +++ b/src/topology/metric_space/basic.lean @@ -54,35 +54,14 @@ open_locale uniformity topology big_operators filter nnreal ennreal universes u v w variables {α : Type u} {β : Type v} {X ι : Type*} -/-- Construct a uniform structure core from a distance function and metric space axioms. -This is a technical construction that can be immediately used to construct a uniform structure -from a distance function and metric space axioms but is also useful when discussing -metrizable topologies, see `pseudo_metric_space.of_metrizable`. -/ -def uniform_space.core_of_dist {α : Type*} (dist : α → α → ℝ) - (dist_self : ∀ x : α, dist x x = 0) - (dist_comm : ∀ x y : α, dist x y = dist y x) - (dist_triangle : ∀ x y z : α, dist x z ≤ dist x y + dist y z) : uniform_space.core α := -{ uniformity := (⨅ ε>0, 𝓟 {p:α×α | dist p.1 p.2 < ε}), - refl := le_infi $ assume ε, le_infi $ - by simp [set.subset_def, id_rel, dist_self, (>)] {contextual := tt}, - comp := le_infi $ assume ε, le_infi $ assume h, lift'_le - (mem_infi_of_mem (ε / 2) $ mem_infi_of_mem (div_pos h zero_lt_two) (subset.refl _)) $ - have ∀ (a b c : α), dist a c < ε / 2 → dist c b < ε / 2 → dist a b < ε, - from assume a b c hac hcb, - calc dist a b ≤ dist a c + dist c b : dist_triangle _ _ _ - ... < ε / 2 + ε / 2 : add_lt_add hac hcb - ... = ε : by rw [div_add_div_same, add_self_div_two], - by simpa [comp_rel], - symm := tendsto_infi.2 $ assume ε, tendsto_infi.2 $ assume h, - tendsto_infi' ε $ tendsto_infi' h $ tendsto_principal_principal.2 $ by simp [dist_comm] } - /-- Construct a uniform structure from a distance function and metric space axioms -/ def uniform_space_of_dist (dist : α → α → ℝ) (dist_self : ∀ x : α, dist x x = 0) (dist_comm : ∀ x y : α, dist x y = dist y x) (dist_triangle : ∀ x y z : α, dist x z ≤ dist x y + dist y z) : uniform_space α := -uniform_space.of_core (uniform_space.core_of_dist dist dist_self dist_comm dist_triangle) +uniform_space.of_fun dist dist_self dist_comm dist_triangle $ λ ε ε0, + ⟨ε / 2, half_pos ε0, λ x hx y hy, add_halves ε ▸ add_lt_add hx hy⟩ /-- This is an internal lemma used to construct a bornology from a metric in `bornology.of_dist`. -/ private lemma bounded_iff_aux {α : Type*} (dist : α → α → ℝ) @@ -202,7 +181,7 @@ instance pseudo_metric_space.to_has_edist : has_edist α := ⟨pseudo_metric_spa /-- Construct a pseudo-metric space structure whose underlying topological space structure (definitionally) agrees which a pre-existing topology which is compatible with a given distance function. -/ -def pseudo_metric_space.of_metrizable {α : Type*} [topological_space α] (dist : α → α → ℝ) +def pseudo_metric_space.of_dist_topology {α : Type u} [topological_space α] (dist : α → α → ℝ) (dist_self : ∀ x : α, dist x x = 0) (dist_comm : ∀ x y : α, dist x y = dist y x) (dist_triangle : ∀ x y z : α, dist x z ≤ dist x y + dist y z) @@ -212,26 +191,11 @@ def pseudo_metric_space.of_metrizable {α : Type*} [topological_space α] (dist dist_self := dist_self, dist_comm := dist_comm, dist_triangle := dist_triangle, - to_uniform_space := { is_open_uniformity := begin - dsimp only [uniform_space.core_of_dist], - intros s, - change is_open s ↔ _, - rw H s, - refine forall₂_congr (λ x x_in, _), - erw (has_basis_binfi_principal _ nonempty_Ioi).mem_iff, - { refine exists₂_congr (λ ε ε_pos, _), - simp only [prod.forall, set_of_subset_set_of], - split, - { rintros h _ y H rfl, - exact h y H }, - { intros h y hxy, - exact h _ _ hxy rfl } }, - { exact λ r (hr : 0 < r) p (hp : 0 < p), ⟨min r p, lt_min hr hp, - λ x (hx : dist _ _ < _), lt_of_lt_of_le hx (min_le_left r p), - λ x (hx : dist _ _ < _), lt_of_lt_of_le hx (min_le_right r p)⟩ }, - { apply_instance } - end, - ..uniform_space.core_of_dist dist dist_self dist_comm dist_triangle }, + to_uniform_space := + { is_open_uniformity := λ s, (H s).trans $ forall₂_congr $ λ x _, + ((uniform_space.has_basis_of_fun (exists_gt (0 : ℝ)) + dist _ _ _ _).comap (prod.mk x)).mem_iff.symm.trans mem_comap_prod_mk, + to_core := (uniform_space_of_dist dist dist_self dist_comm dist_triangle).to_core }, uniformity_dist := rfl, to_bornology := bornology.of_dist dist dist_self dist_comm dist_triangle, cobounded_sets := rfl } @@ -653,14 +617,15 @@ theorem is_bounded_iff_nndist {s : set α} : by simp only [is_bounded_iff_exists_ge 0, nnreal.exists, ← nnreal.coe_le_coe, ← dist_nndist, nnreal.coe_mk, exists_prop] +theorem to_uniform_space_eq : ‹pseudo_metric_space α›.to_uniform_space = + uniform_space_of_dist dist dist_self dist_comm dist_triangle := +uniform_space_eq pseudo_metric_space.uniformity_dist + theorem uniformity_basis_dist : (𝓤 α).has_basis (λ ε : ℝ, 0 < ε) (λ ε, {p:α×α | dist p.1 p.2 < ε}) := begin - rw ← pseudo_metric_space.uniformity_dist.symm, - refine has_basis_binfi_principal _ nonempty_Ioi, - exact λ r (hr : 0 < r) p (hp : 0 < p), ⟨min r p, lt_min hr hp, - λ x (hx : dist _ _ < _), lt_of_lt_of_le hx (min_le_left r p), - λ x (hx : dist _ _ < _), lt_of_lt_of_le hx (min_le_right r p)⟩ + rw [to_uniform_space_eq], + exact uniform_space.has_basis_of_fun (exists_gt _) _ _ _ _ _ end /-- Given `f : β → ℝ`, if `f` sends `{i | p i}` to a set of positive numbers @@ -680,6 +645,11 @@ begin { exact λ ⟨i, hi, H⟩, ⟨f i, hf₀ i hi, H⟩ } end +theorem uniformity_basis_dist_rat : + (𝓤 α).has_basis (λ r : ℚ, 0 < r) (λ r, {p : α × α | dist p.1 p.2 < r}) := +metric.mk_uniformity_basis (λ _, rat.cast_pos.2) $ λ ε hε, + let ⟨r, hr0, hrε⟩ := exists_rat_btwn hε in ⟨r, rat.cast_pos.1 hr0, hrε.le⟩ + theorem uniformity_basis_dist_inv_nat_succ : (𝓤 α).has_basis (λ _, true) (λ n:ℕ, {p:α×α | dist p.1 p.2 < 1 / (↑n+1) }) := metric.mk_uniformity_basis (λ n _, div_pos zero_lt_one $ nat.cast_add_one_pos n) @@ -1476,11 +1446,7 @@ def pseudo_metric_space.induced {α β} (f : α → β) edist := λ x y, edist (f x) (f y), edist_dist := λ x y, edist_dist _ _, to_uniform_space := uniform_space.comap f m.to_uniform_space, - uniformity_dist := begin - apply @uniformity_dist_of_mem_uniformity _ _ _ _ _ (λ x y, dist (f x) (f y)), - refine compl_surjective.forall.2 (λ s, compl_mem_comap.trans $ mem_uniformity_dist.trans _), - simp only [mem_compl_iff, @imp_not_comm _ (_ ∈ _), ← prod.forall', prod.mk.eta, ball_image_iff] - end, + uniformity_dist := (uniformity_basis_dist.comap _).eq_binfi, to_bornology := bornology.induced f, cobounded_sets := set.ext $ compl_surjective.forall.2 $ λ s, by simp only [compl_mem_comap, filter.mem_sets, ← is_bounded_def, mem_set_of_eq, compl_compl, @@ -2643,14 +2609,14 @@ end /-- Construct a metric space structure whose underlying topological space structure (definitionally) agrees which a pre-existing topology which is compatible with a given distance function. -/ -def metric_space.of_metrizable {α : Type*} [topological_space α] (dist : α → α → ℝ) +def metric_space.of_dist_topology {α : Type u} [topological_space α] (dist : α → α → ℝ) (dist_self : ∀ x : α, dist x x = 0) (dist_comm : ∀ x y : α, dist x y = dist y x) (dist_triangle : ∀ x y z : α, dist x z ≤ dist x y + dist y z) (H : ∀ s : set α, is_open s ↔ ∀ x ∈ s, ∃ ε > 0, ∀ y, dist x y < ε → y ∈ s) (eq_of_dist_eq_zero : ∀ x y : α, dist x y = 0 → x = y) : metric_space α := { eq_of_dist_eq_zero := eq_of_dist_eq_zero, - ..pseudo_metric_space.of_metrizable dist dist_self dist_comm dist_triangle H } + ..pseudo_metric_space.of_dist_topology dist dist_self dist_comm dist_triangle H } variables {γ : Type w} [metric_space γ] diff --git a/src/topology/metric_space/emetric_space.lean b/src/topology/metric_space/emetric_space.lean index 64cde1bbe17d1..c28d6544f25e9 100644 --- a/src/topology/metric_space/emetric_space.lean +++ b/src/topology/metric_space/emetric_space.lean @@ -38,41 +38,20 @@ variables {α : Type u} {β : Type v} {X : Type*} in terms of the elements of the uniformity. -/ theorem uniformity_dist_of_mem_uniformity [linear_order β] {U : filter (α × α)} (z : β) (D : α → α → β) (H : ∀ s, s ∈ U ↔ ∃ε>z, ∀{a b:α}, D a b < ε → (a, b) ∈ s) : - U = ⨅ ε>z, 𝓟 {p:α×α | D p.1 p.2 < ε} := -le_antisymm - (le_infi $ λ ε, le_infi $ λ ε0, le_principal_iff.2 $ (H _).2 ⟨ε, ε0, λ a b, id⟩) - (λ r ur, let ⟨ε, ε0, h⟩ := (H _).1 ur in - mem_infi_of_mem ε $ mem_infi_of_mem ε0 $ mem_principal.2 $ λ ⟨a, b⟩, h) + U = ⨅ ε > z, 𝓟 {p:α×α | D p.1 p.2 < ε} := +has_basis.eq_binfi ⟨λ s, by simp only [H, subset_def, prod.forall, mem_set_of]⟩ /-- `has_edist α` means that `α` is equipped with an extended distance. -/ class has_edist (α : Type*) := (edist : α → α → ℝ≥0∞) export has_edist (edist) /-- Creating a uniform space from an extended distance. -/ -def uniform_space_of_edist - (edist : α → α → ℝ≥0∞) - (edist_self : ∀ x : α, edist x x = 0) - (edist_comm : ∀ x y : α, edist x y = edist y x) +noncomputable def uniform_space_of_edist (edist : α → α → ℝ≥0∞) + (edist_self : ∀ x : α, edist x x = 0) (edist_comm : ∀ x y : α, edist x y = edist y x) (edist_triangle : ∀ x y z : α, edist x z ≤ edist x y + edist y z) : uniform_space α := -uniform_space.of_core -{ uniformity := (⨅ ε>0, 𝓟 {p:α×α | edist p.1 p.2 < ε}), - refl := le_infi $ assume ε, le_infi $ - by simp [set.subset_def, id_rel, edist_self, (>)] {contextual := tt}, - comp := - le_infi $ assume ε, le_infi $ assume h, - have (2 : ℝ≥0∞) = (2 : ℕ) := by simp, - have A : 0 < ε / 2 := ennreal.div_pos_iff.2 - ⟨ne_of_gt h, by { convert ennreal.nat_ne_top 2 }⟩, - lift'_le - (mem_infi_of_mem (ε / 2) $ mem_infi_of_mem A (subset.refl _)) $ - have ∀ (a b c : α), edist a c < ε / 2 → edist c b < ε / 2 → edist a b < ε, - from assume a b c hac hcb, - calc edist a b ≤ edist a c + edist c b : edist_triangle _ _ _ - ... < ε / 2 + ε / 2 : ennreal.add_lt_add hac hcb - ... = ε : by rw [ennreal.add_halves], - by simpa [comp_rel], - symm := tendsto_infi.2 $ assume ε, tendsto_infi.2 $ assume h, - tendsto_infi' ε $ tendsto_infi' h $ tendsto_principal_principal.2 $ by simp [edist_comm] } +uniform_space.of_fun edist edist_self edist_comm edist_triangle $ + λ ε ε0, ⟨ε / 2, ennreal.half_pos ε0.lt.ne', λ _ h₁ _ h₂, + (ennreal.add_lt_add h₁ h₂).trans_eq (ennreal.add_halves _)⟩ -- the uniform structure is embedded in the emetric space structure -- to avoid instance diamond issues. See Note [forgetful inheritance]. @@ -175,13 +154,13 @@ theorem uniformity_pseudoedist : 𝓤 α = ⨅ ε>0, 𝓟 {p:α×α | edist p.1 p.2 < ε} := pseudo_emetric_space.uniformity_edist +theorem uniform_space_edist : ‹pseudo_emetric_space α›.to_uniform_space = + uniform_space_of_edist edist edist_self edist_comm edist_triangle := +uniform_space_eq uniformity_pseudoedist + theorem uniformity_basis_edist : (𝓤 α).has_basis (λ ε : ℝ≥0∞, 0 < ε) (λ ε, {p:α×α | edist p.1 p.2 < ε}) := -(@uniformity_pseudoedist α _).symm ▸ has_basis_binfi_principal - (λ r hr p hp, ⟨min r p, lt_min hr hp, - λ x hx, lt_of_lt_of_le hx (min_le_left _ _), - λ x hx, lt_of_lt_of_le hx (min_le_right _ _)⟩) - ⟨1, zero_lt_one⟩ +(@uniform_space_edist α _).symm ▸ uniform_space.has_basis_of_fun ⟨1, one_pos⟩ _ _ _ _ _ /-- Characterization of the elements of the uniformity in terms of the extended distance -/ theorem mem_uniformity_edist {s : set (α×α)} : @@ -394,16 +373,7 @@ def pseudo_emetric_space.induced {α β} (f : α → β) edist_comm := λ x y, edist_comm _ _, edist_triangle := λ x y z, edist_triangle _ _ _, to_uniform_space := uniform_space.comap f m.to_uniform_space, - uniformity_edist := begin - apply @uniformity_dist_of_mem_uniformity _ _ _ _ _ (λ x y, edist (f x) (f y)), - refine λ s, mem_comap.trans _, - split; intro H, - { rcases H with ⟨r, ru, rs⟩, - rcases mem_uniformity_edist.1 ru with ⟨ε, ε0, hε⟩, - refine ⟨ε, ε0, λ a b h, rs (hε _)⟩, exact h }, - { rcases H with ⟨ε, ε0, hε⟩, - exact ⟨_, edist_mem_uniformity ε0, λ ⟨a, b⟩, hε⟩ } - end } + uniformity_edist := (uniformity_basis_edist.comap _).eq_binfi } /-- Pseudoemetric space instance on subsets of pseudoemetric spaces -/ instance {α : Type*} {p : α → Prop} [pseudo_emetric_space α] : pseudo_emetric_space (subtype p) := @@ -978,16 +948,7 @@ def emetric_space.induced {γ β} (f : γ → β) (hf : function.injective f) edist_comm := λ x y, edist_comm _ _, edist_triangle := λ x y z, edist_triangle _ _ _, to_uniform_space := uniform_space.comap f m.to_uniform_space, - uniformity_edist := begin - apply @uniformity_dist_of_mem_uniformity _ _ _ _ _ (λ x y, edist (f x) (f y)), - refine λ s, mem_comap.trans _, - split; intro H, - { rcases H with ⟨r, ru, rs⟩, - rcases mem_uniformity_edist.1 ru with ⟨ε, ε0, hε⟩, - refine ⟨ε, ε0, λ a b h, rs (hε _)⟩, exact h }, - { rcases H with ⟨ε, ε0, hε⟩, - exact ⟨_, edist_mem_uniformity ε0, λ ⟨a, b⟩, hε⟩ } - end } + uniformity_edist := (uniformity_basis_edist.comap _).eq_binfi } /-- Emetric space instance on subsets of emetric spaces -/ instance {α : Type*} {p : α → Prop} [emetric_space α] : emetric_space (subtype p) := diff --git a/src/topology/metric_space/gluing.lean b/src/topology/metric_space/gluing.lean index da4d50cd2a0dd..9071591df748d 100644 --- a/src/topology/metric_space/gluing.lean +++ b/src/topology/metric_space/gluing.lean @@ -464,7 +464,7 @@ their respective basepoints, plus the distance 1 between the basepoints. Since there is an arbitrary choice in this construction, it is not an instance by default. -/ protected def metric_space : metric_space (Σ i, E i) := begin - refine metric_space.of_metrizable sigma.dist _ _ sigma.dist_triangle + refine metric_space.of_dist_topology sigma.dist _ _ sigma.dist_triangle sigma.is_open_iff _, { rintros ⟨i, x⟩, simp [sigma.dist] }, { rintros ⟨i, x⟩ ⟨j, y⟩, diff --git a/src/topology/metric_space/pi_nat.lean b/src/topology/metric_space/pi_nat.lean index 1038852859a4c..d1ff39260c202 100644 --- a/src/topology/metric_space/pi_nat.lean +++ b/src/topology/metric_space/pi_nat.lean @@ -357,7 +357,7 @@ but it does not take care of a possible uniformity. If the `E n` have a uniform there will be two non-defeq uniform structures on `Π n, E n`, the product one and the one coming from the metric structure. In this case, use `metric_space_of_discrete_uniformity` instead. -/ protected def metric_space : metric_space (Π n, E n) := -metric_space.of_metrizable dist pi_nat.dist_self pi_nat.dist_comm pi_nat.dist_triangle +metric_space.of_dist_topology dist pi_nat.dist_self pi_nat.dist_comm pi_nat.dist_triangle is_open_iff_dist pi_nat.eq_of_dist_eq_zero /-- Metric space structure on `Π (n : ℕ), E n` when the spaces `E n` have the discrete uniformity, diff --git a/src/topology/uniform_space/absolute_value.lean b/src/topology/uniform_space/absolute_value.lean index 5796516c80e8e..e4640629275f9 100644 --- a/src/topology/uniform_space/absolute_value.lean +++ b/src/topology/uniform_space/absolute_value.lean @@ -33,46 +33,21 @@ absolute value, uniform spaces -/ open set function filter uniform_space -open_locale filter +open_locale filter topology -namespace is_absolute_value -variables {𝕜 : Type*} [linear_ordered_field 𝕜] -variables {R : Type*} [comm_ring R] (abv : R → 𝕜) [is_absolute_value abv] +namespace absolute_value -/-- The uniformity coming from an absolute value. -/ -def uniform_space_core : uniform_space.core R := -{ uniformity := (⨅ ε>0, 𝓟 {p:R×R | abv (p.2 - p.1) < ε}), - refl := le_infi $ assume ε, le_infi $ assume ε_pos, principal_mono.2 - (λ ⟨x, y⟩ h, by simpa [show x = y, from h, abv_zero abv]), - symm := tendsto_infi.2 $ assume ε, tendsto_infi.2 $ assume h, - tendsto_infi' ε $ tendsto_infi' h $ tendsto_principal_principal.2 $ λ ⟨x, y⟩ h, - have h : abv (y - x) < ε, by simpa [-sub_eq_add_neg] using h, - by rwa abv_sub abv at h, - comp := le_infi $ assume ε, le_infi $ assume h, lift'_le - (mem_infi_of_mem (ε / 2) $ mem_infi_of_mem (div_pos h zero_lt_two) (subset.refl _)) $ - have ∀ (a b c : R), abv (c-a) < ε / 2 → abv (b-c) < ε / 2 → abv (b-a) < ε, - from assume a b c hac hcb, - calc abv (b - a) ≤ _ : abv_sub_le abv b c a - ... = abv (c - a) + abv (b - c) : add_comm _ _ - ... < ε / 2 + ε / 2 : add_lt_add hac hcb - ... = ε : by rw [div_add_div_same, add_self_div_two], - by simpa [comp_rel] } +variables {𝕜 : Type*} [linear_ordered_field 𝕜] +variables {R : Type*} [comm_ring R] (abv : absolute_value R 𝕜) -/-- The uniform structure coming from an absolute value. -/ -def uniform_space : uniform_space R := -uniform_space.of_core (uniform_space_core abv) +/-- The uniform space structure coming from an absolute value. -/ +protected def uniform_space : uniform_space R := +uniform_space.of_fun (λ x y, abv (y - x)) (by simp) (λ x y, abv.map_sub y x) + (λ x y z, (abv.sub_le _ _ _).trans_eq (add_comm _ _)) $ + λ ε ε0, ⟨ε / 2, half_pos ε0, λ _ h₁ _ h₂, (add_lt_add h₁ h₂).trans_eq (add_halves ε)⟩ -theorem mem_uniformity {s : set (R×R)} : - s ∈ (uniform_space_core abv).uniformity ↔ - (∃ε>0, ∀{a b:R}, abv (b - a) < ε → (a, b) ∈ s) := -begin - suffices : s ∈ (⨅ ε: {ε : 𝕜 // ε > 0}, 𝓟 {p:R×R | abv (p.2 - p.1) < ε.val}) ↔ _, - { rw infi_subtype at this, - exact this }, - rw mem_infi_of_directed, - { simp [subset_def] }, - { rintros ⟨r, hr⟩ ⟨p, hp⟩, - exact ⟨⟨min r p, lt_min hr hp⟩, by simp [lt_min_iff, (≥)] {contextual := tt}⟩, }, -end +theorem has_basis_uniformity : + 𝓤[abv.uniform_space].has_basis (λ ε : 𝕜, 0 < ε) (λ ε, {p : R × R | abv (p.2 - p.1) < ε}) := +uniform_space.has_basis_of_fun (exists_gt _) _ _ _ _ _ -end is_absolute_value +end absolute_value diff --git a/src/topology/uniform_space/basic.lean b/src/topology/uniform_space/basic.lean index 178ffdb747e55..a7ef9b0de87c2 100644 --- a/src/topology/uniform_space/basic.lean +++ b/src/topology/uniform_space/basic.lean @@ -316,6 +316,32 @@ lemma uniform_space.replace_topology_eq {α : Type*} [i : topological_space α] (h : i = u.to_topological_space) : u.replace_topology h = u := u.of_core_eq_to_core _ _ +/-- Define a `uniform_space` using a "distance" function. The function can be, e.g., the distance in +a (usual or extended) metric space or an absolute value on a ring. -/ +def uniform_space.of_fun {α β : Type*} [ordered_add_comm_monoid β] + (d : α → α → β) (refl : ∀ x, d x x = 0) (symm : ∀ x y, d x y = d y x) + (triangle : ∀ x y z, d x z ≤ d x y + d y z) + (half : ∀ ε > (0 : β), ∃ δ > (0 : β), ∀ x < δ, ∀ y < δ, x + y < ε) : + uniform_space α := +uniform_space.of_core + { uniformity := ⨅ r > 0, 𝓟 { x | d x.1 x.2 < r }, + refl := le_infi₂ $ λ r hr, principal_mono.2 $ id_rel_subset.2 $ λ x, by simpa [refl], + symm := tendsto_infi_infi $ λ r, tendsto_infi_infi $ λ _, tendsto_principal_principal.2 $ + λ x hx, by rwa [mem_set_of, symm], + comp := le_infi₂ $ λ r hr, let ⟨δ, h0, hδr⟩ := half r hr in le_principal_iff.2 $ mem_of_superset + (mem_lift' $ mem_infi_of_mem δ $ mem_infi_of_mem h0 $ mem_principal_self _) $ + λ ⟨x, z⟩ ⟨y, h₁, h₂⟩, (triangle _ _ _).trans_lt (hδr _ h₁ _ h₂) } + +lemma uniform_space.has_basis_of_fun {α β : Type*} [linear_ordered_add_comm_monoid β] + (h₀ : ∃ x : β, 0 < x) (d : α → α → β) (refl : ∀ x, d x x = 0) (symm : ∀ x y, d x y = d y x) + (triangle : ∀ x y z, d x z ≤ d x y + d y z) + (half : ∀ ε > (0 : β), ∃ δ > (0 : β), ∀ x < δ, ∀ y < δ, x + y < ε) : + 𝓤[uniform_space.of_fun d refl symm triangle half].has_basis ((<) (0 : β)) + (λ ε, { x | d x.1 x.2 < ε }) := +has_basis_binfi_principal' + (λ ε₁ h₁ ε₂ h₂, ⟨min ε₁ ε₂, lt_min h₁ h₂, λ _x hx, lt_of_lt_of_le hx (min_le_left _ _), + λ _x hx, lt_of_lt_of_le hx (min_le_right _ _)⟩) h₀ + section uniform_space variables [uniform_space α] diff --git a/src/topology/uniform_space/compare_reals.lean b/src/topology/uniform_space/compare_reals.lean index 4da0b8014a4e1..c8fd75c649d96 100644 --- a/src/topology/uniform_space/compare_reals.lean +++ b/src/topology/uniform_space/compare_reals.lean @@ -55,27 +55,16 @@ open set function filter cau_seq uniform_space /-- The metric space uniform structure on ℚ (which presupposes the existence of real numbers) agrees with the one coming directly from (abs : ℚ → ℚ). -/ lemma rat.uniform_space_eq : - is_absolute_value.uniform_space (abs : ℚ → ℚ) = pseudo_metric_space.to_uniform_space := + (absolute_value.abs : absolute_value ℚ ℚ).uniform_space = pseudo_metric_space.to_uniform_space := begin ext s, - erw [metric.mem_uniformity_dist, is_absolute_value.mem_uniformity], - split ; rintro ⟨ε, ε_pos, h⟩, - { use [ε, by exact_mod_cast ε_pos], - intros a b hab, - apply h, - rw [rat.dist_eq, abs_sub_comm] at hab, - exact_mod_cast hab }, - { obtain ⟨ε', h', h''⟩ : ∃ ε' : ℚ, 0 < ε' ∧ (ε' : ℝ) < ε, from exists_pos_rat_lt ε_pos, - use [ε', h'], - intros a b hab, - apply h, - rw [rat.dist_eq, abs_sub_comm], - refine lt_trans _ h'', - exact_mod_cast hab } + rw [(absolute_value.has_basis_uniformity _).mem_iff, metric.uniformity_basis_dist_rat.mem_iff], + simp only [rat.dist_eq, absolute_value.abs_apply, ← rat.cast_sub, ← rat.cast_abs, rat.cast_lt, + abs_sub_comm] end /-- Cauchy reals packaged as a completion of ℚ using the absolute value route. -/ -def rational_cau_seq_pkg : @abstract_completion ℚ $ is_absolute_value.uniform_space (abs : ℚ → ℚ) := +def rational_cau_seq_pkg : @abstract_completion ℚ $ (@absolute_value.abs ℚ _).uniform_space := { space := ℝ, coe := (coe : ℚ → ℝ), uniform_struct := by apply_instance, @@ -92,7 +81,7 @@ but they are not definitionaly equal, so it would confuse the type class system also human readers). -/ @[derive comm_ring, derive inhabited] def Q := ℚ -instance : uniform_space Q := is_absolute_value.uniform_space (abs : ℚ → ℚ) +instance : uniform_space Q := (@absolute_value.abs ℚ _).uniform_space /-- Real numbers constructed as in Bourbaki. -/ @[derive inhabited] From 87ecf9614fedda186a792c38ee571904f548df29 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Tue, 28 Feb 2023 01:51:59 +0000 Subject: [PATCH 0554/1291] chore(geometry/manifold/vector_bundle): golf (#18509) Golf the workaround introduced in #18507 --- src/geometry/manifold/vector_bundle.lean | 49 +++--------------------- 1 file changed, 5 insertions(+), 44 deletions(-) diff --git a/src/geometry/manifold/vector_bundle.lean b/src/geometry/manifold/vector_bundle.lean index edc4850291457..1a74b0de7f3b0 100644 --- a/src/geometry/manifold/vector_bundle.lean +++ b/src/geometry/manifold/vector_bundle.lean @@ -120,52 +120,13 @@ lemma smooth_fiberwise_linear.locality_aux₁ (e : local_homeomorph (B × F) (B (e.restr (u ×ˢ univ)).eq_on_source (fiberwise_linear.local_homeomorph φ hu hφ.continuous_on h2φ.continuous_on) := begin - rw set_coe.forall' at h, + rw [set_coe.forall'] at h, -- choose s hs hsp φ u hu hφ h2φ heφ using h, - -- the following ~40 lines should be `choose s hs hsp φ u hu hφ h2φ heφ using h,` + -- the following 2 lines should be `choose s hs hsp φ u hu hφ h2φ heφ using h,` -- `choose` produces a proof term that takes a long time to type-check by the kernel (it seems) - let s := λ x, classical.some (h x), - have hs : ∀ x, is_open (s x) := λ x, (classical.some_spec (h x)).1, - have hsp : ∀ x : e.source, x.1 ∈ s x := λ x, (classical.some_spec (h x)).2.1, - have h : ∀ x : e.source, ∃ (φ : B → (F ≃L[𝕜] F)) (u : set B) (hu : is_open u) - (hφ : smooth_on IB 𝓘(𝕜, F →L[𝕜] F) (λ x, (φ x : F →L[𝕜] F)) u) - (h2φ : smooth_on IB 𝓘(𝕜, F →L[𝕜] F) (λ x, ((φ x).symm : F →L[𝕜] F)) u), - (e.restr (s x)).eq_on_source - (fiberwise_linear.local_homeomorph φ hu hφ.continuous_on h2φ.continuous_on) := - λ x, (classical.some_spec (h x)).2.2, - let φ : e.source → B → (F ≃L[𝕜] F) := λ x, classical.some (h x), - have h : ∀ z : e.source, ∃ (u : set B) (hu : is_open u) - (hφ : smooth_on IB 𝓘(𝕜, F →L[𝕜] F) (λ x, (φ z x : F →L[𝕜] F)) u) - (h2φ : smooth_on IB 𝓘(𝕜, F →L[𝕜] F) (λ x, ((φ z x).symm : F →L[𝕜] F)) u), - (e.restr (s z)).eq_on_source - (fiberwise_linear.local_homeomorph (φ z) hu hφ.continuous_on h2φ.continuous_on) := - λ x, classical.some_spec (h x), - let u : e.source → set B := λ x, classical.some (h x), - have h : ∀ z : e.source, ∃ (hu : is_open (u z)) - (hφ : smooth_on IB 𝓘(𝕜, F →L[𝕜] F) (λ x, (φ z x : F →L[𝕜] F)) (u z)) - (h2φ : smooth_on IB 𝓘(𝕜, F →L[𝕜] F) (λ x, ((φ z x).symm : F →L[𝕜] F)) (u z)), - (e.restr (s z)).eq_on_source - (fiberwise_linear.local_homeomorph (φ z) hu hφ.continuous_on h2φ.continuous_on) := - λ x, classical.some_spec (h x), - have hu : ∀ x, is_open (u x) := λ x, classical.some (h x), - have h : ∀ z : e.source, ∃ (hφ : smooth_on IB 𝓘(𝕜, F →L[𝕜] F) (λ x, (φ z x : F →L[𝕜] F)) (u z)) - (h2φ : smooth_on IB 𝓘(𝕜, F →L[𝕜] F) (λ x, ((φ z x).symm : F →L[𝕜] F)) (u z)), - (e.restr (s z)).eq_on_source - (fiberwise_linear.local_homeomorph (φ z) (hu z) hφ.continuous_on h2φ.continuous_on) := - λ x, classical.some_spec (h x), - have hφ : ∀ z, smooth_on IB 𝓘(𝕜, F →L[𝕜] F) (λ x, (φ z x : F →L[𝕜] F)) (u z) := - λ x, classical.some (h x), - have h : ∀ z : e.source, - ∃ (h2φ : smooth_on IB 𝓘(𝕜, F →L[𝕜] F) (λ x, ((φ z x).symm : F →L[𝕜] F)) (u z)), - (e.restr (s z)).eq_on_source - (fiberwise_linear.local_homeomorph (φ z) (hu z) (hφ z).continuous_on h2φ.continuous_on) := - λ x, classical.some_spec (h x), - have h2φ : ∀ z, smooth_on IB 𝓘(𝕜, F →L[𝕜] F) (λ x, ((φ z x).symm : F →L[𝕜] F)) (u z) := - λ x, classical.some (h x), - have heφ : ∀ z, (e.restr (s z)).eq_on_source - (fiberwise_linear.local_homeomorph (φ z) (hu z) (hφ z).continuous_on (h2φ z).continuous_on) := - λ x, classical.some_spec (h x), - clear_value s φ u, clear h h h h h h, + -- porting note: todo: try using `choose` again in Lean 4 + simp only [classical.skolem, ← exists_prop] at h, + rcases h with ⟨s, hs, hsp, φ, u, hu, hφ, h2φ, heφ⟩, have hesu : ∀ p : e.source, e.source ∩ s p = u p ×ˢ univ, { intros p, rw ← e.restr_source' (s _) (hs _), From 7b1d4abc172ea1bcf200d11041d61a4fba058023 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Tue, 28 Feb 2023 04:45:17 +0000 Subject: [PATCH 0555/1291] feat(order/game_add): `sym2.game_add` (#18287) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We define `sym2.game_add`, which is a relation `sym2 α → sym2 α → Prop` which expresses that one pair may be obtained from the other by decreasing either entry. We add a basic API, and prove well-foundedness. Pair: https://github.com/leanprover-community/mathlib4/pull/2532 --- src/order/game_add.lean | 103 ++++++++++++++++++++++++++++++++++++---- 1 file changed, 95 insertions(+), 8 deletions(-) diff --git a/src/order/game_add.lean b/src/order/game_add.lean index d8e8a15a9f702..6534e70f81c8e 100644 --- a/src/order/game_add.lean +++ b/src/order/game_add.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Junyan Xu. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Junyan Xu -/ -import order.basic +import data.sym.sym2 import logic.relation /-! @@ -23,15 +23,19 @@ subsequency relation on the addition of combinatorial games. - `well_founded.prod_game_add`: formalizes induction on ordered pairs, where exactly one entry decreases at a time. -## Todo - -- Define `sym2.game_add`. +- `sym2.game_add`: the game addition relation on unordered pairs. +- `well_founded.sym2_game_add`: formalizes induction on unordered pairs, where exactly one entry + decreases at a time. -/ -variables {α β : Type*} (rα : α → α → Prop) (rβ : β → β → Prop) +variables {α β : Type*} {rα : α → α → Prop} {rβ : β → β → Prop} + +/-! ### `prod.game_add` -/ namespace prod +variables (rα rβ) + /-- `prod.game_add rα rβ x y` means that `x` can be reached from `y` by decreasing either entry with respect to the relations `rα` and `rβ`. @@ -63,6 +67,10 @@ game_add_iff game_add rβ rα a.swap b.swap ↔ game_add rα rβ a b := λ ⟨a₁, b₁⟩ ⟨a₂, b₂⟩, by rw [prod.swap, game_add_mk_iff, game_add_mk_iff, or_comm] +lemma game_add_swap_swap_mk (a₁ a₂ : α) (b₁ b₂ : β) : + game_add rα rβ (a₁, b₁) (a₂, b₂) ↔ game_add rβ rα (b₁, a₁) (b₂, a₂) := +game_add_swap_swap rβ rα (b₁, a₁) (b₂, a₂) + /-- `prod.game_add` is a `subrelation` of `prod.lex`. -/ lemma game_add_le_lex : game_add rα rβ ≤ prod.lex rα rβ := λ _ _ h, h.rec (λ _ _ b, prod.lex.left b b) (λ a _ _, prod.lex.right a) @@ -76,8 +84,6 @@ end end prod -variables {rα rβ} - /-- If `a` is accessible under `rα` and `b` is accessible under `rβ`, then `(a, b)` is accessible under `prod.game_add rα rβ`. Notice that `prod.lex_accessible` requires the stronger condition `∀ b, acc rβ b`. -/ @@ -110,7 +116,7 @@ def game_add.fix {C : α → β → Sort*} (hα : well_founded rα) (hβ : well_ lemma game_add.fix_eq {C : α → β → Sort*} (hα : well_founded rα) (hβ : well_founded rβ) (IH : Π a₁ b₁, (Π a₂ b₂, game_add rα rβ (a₂, b₂) (a₁, b₁) → C a₂ b₂) → C a₁ b₁) (a : α) (b : β) : game_add.fix hα hβ IH a b = IH a b (λ a' b' h, game_add.fix hα hβ IH a' b') := -by { rw [game_add.fix, well_founded.fix_eq], refl } +well_founded.fix_eq _ _ _ /-- Induction on the well-founded `prod.game_add` relation. @@ -120,3 +126,84 @@ lemma game_add.induction {C : α → β → Prop} : well_founded rα → well_fo game_add.fix end prod + +/-! ### `sym2.game_add` -/ + +namespace sym2 + +/-- `sym2.game_add rα x y` means that `x` can be reached from `y` by decreasing either entry. -/ +def game_add (rα : α → α → Prop): sym2 α → sym2 α → Prop := +sym2.lift₂ +⟨λ a₁ b₁ a₂ b₂, prod.game_add rα rα (a₁, b₁) (a₂, b₂) ∨ prod.game_add rα rα (b₁, a₁) (a₂, b₂), + λ a₁ b₁ a₂ b₂, begin + rw [prod.game_add_swap_swap_mk _ _ b₁ b₂ a₁ a₂, prod.game_add_swap_swap_mk _ _ a₁ b₂ b₁ a₂], + simp [or_comm] + end⟩ + +variable {rα} + +lemma game_add_iff : ∀ {x y : α × α}, game_add rα ⟦x⟧ ⟦y⟧ ↔ + prod.game_add rα rα x y ∨ prod.game_add rα rα x.swap y := +by { rintros ⟨_, _⟩ ⟨_, _⟩, refl } + +lemma game_add_mk_iff {a₁ a₂ b₁ b₂ : α} : game_add rα ⟦(a₁, b₁)⟧ ⟦(a₂, b₂)⟧ ↔ + prod.game_add rα rα (a₁, b₁) (a₂, b₂) ∨ prod.game_add rα rα (b₁, a₁) (a₂, b₂) := +iff.rfl + +lemma _root_.prod.game_add.to_sym2 {a₁ a₂ b₁ b₂ : α} + (h : prod.game_add rα rα (a₁, b₁) (a₂, b₂)) : sym2.game_add rα ⟦(a₁, b₁)⟧ ⟦(a₂, b₂)⟧ := +game_add_mk_iff.2 $ or.inl $ h + +lemma game_add.fst {a₁ a₂ b : α} (h : rα a₁ a₂) : game_add rα ⟦(a₁, b)⟧ ⟦(a₂, b)⟧ := +(prod.game_add.fst h).to_sym2 + +lemma game_add.snd {a b₁ b₂ : α} (h : rα b₁ b₂) : game_add rα ⟦(a, b₁)⟧ ⟦(a, b₂)⟧ := +(prod.game_add.snd h).to_sym2 + +lemma game_add.fst_snd {a₁ a₂ b : α} (h : rα a₁ a₂) : game_add rα ⟦(a₁, b)⟧ ⟦(b, a₂)⟧ := +by { rw sym2.eq_swap, exact game_add.snd h } + +lemma game_add.snd_fst {a₁ a₂ b : α} (h : rα a₁ a₂) : game_add rα ⟦(b, a₁)⟧ ⟦(a₂, b)⟧ := +by { rw sym2.eq_swap, exact game_add.fst h } + +end sym2 + +lemma acc.sym2_game_add {a b} (ha : acc rα a) (hb : acc rα b) : acc (sym2.game_add rα) ⟦(a, b)⟧ := +begin + induction ha with a ha iha generalizing b, + induction hb with b hb ihb, + refine acc.intro _ (λ s, _), + induction s using sym2.induction_on with c d, + rintros ((rc | rd) | (rd | rc)), + { exact iha c rc ⟨b, hb⟩ }, + { exact ihb d rd }, + { rw sym2.eq_swap, + exact iha d rd ⟨b, hb⟩ }, + { rw sym2.eq_swap, + exact ihb c rc } +end + +/-- The `sym2.game_add` relation on well-founded inputs is well-founded. -/ +lemma well_founded.sym2_game_add (h : well_founded rα) : well_founded (sym2.game_add rα) := +⟨λ i, sym2.induction_on i $ λ x y, (h.apply x).sym2_game_add (h.apply y)⟩ + +namespace sym2 + +/-- Recursion on the well-founded `sym2.game_add` relation. -/ +def game_add.fix {C : α → α → Sort*} (hr : well_founded rα) + (IH : Π a₁ b₁, (Π a₂ b₂, sym2.game_add rα ⟦(a₂, b₂)⟧ ⟦(a₁, b₁)⟧ → C a₂ b₂) → C a₁ b₁) (a b : α) : + C a b := +@well_founded.fix (α × α) (λ x, C x.1 x.2) _ hr.sym2_game_add.of_quotient_lift₂ + (λ ⟨x₁, x₂⟩ IH', IH x₁ x₂ $ λ a' b', IH' ⟨a', b'⟩) (a, b) + +lemma game_add.fix_eq {C : α → α → Sort*} (hr : well_founded rα) + (IH : Π a₁ b₁, (Π a₂ b₂, sym2.game_add rα ⟦(a₂, b₂)⟧ ⟦(a₁, b₁)⟧ → C a₂ b₂) → C a₁ b₁) (a b : α) : + game_add.fix hr IH a b = IH a b (λ a' b' h, game_add.fix hr IH a' b') := +well_founded.fix_eq _ _ _ + +/-- Induction on the well-founded `sym2.game_add` relation. -/ +lemma game_add.induction {C : α → α → Prop} : well_founded rα → + (∀ a₁ b₁, (∀ a₂ b₂, sym2.game_add rα ⟦(a₂, b₂)⟧ ⟦(a₁, b₁)⟧ → C a₂ b₂) → C a₁ b₁) → ∀ a b, C a b := +game_add.fix + +end sym2 From d7feae342946021379fc1e5025856afbce1de5dd Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Tue, 28 Feb 2023 09:17:58 +0000 Subject: [PATCH 0556/1291] feat(set_theory/zfc/basic): induction principle for sets (#18324) If every subset of a class is a member, the class is universal. This will help us establish that the von Neumann hierarchy exhausts sets later on. --- src/set_theory/zfc/basic.lean | 12 ++++++++++++ 1 file changed, 12 insertions(+) diff --git a/src/set_theory/zfc/basic.lean b/src/set_theory/zfc/basic.lean index 7396b9188f529..ce57ca159708e 100644 --- a/src/set_theory/zfc/basic.lean +++ b/src/set_theory/zfc/basic.lean @@ -927,6 +927,9 @@ exists_congr $ λ x, and_true _ @[simp] theorem mem_univ_hom (x : Set.{u}) : univ.{u} x := trivial +theorem eq_univ_iff_forall {A : Class.{u}} : A = univ ↔ ∀ x : Set, A x := set.eq_univ_iff_forall +theorem eq_univ_of_forall {A : Class.{u}} : (∀ x : Set, A x) → A = univ := set.eq_univ_of_forall + theorem mem_wf : @well_founded Class.{u} (∈) := ⟨begin have H : ∀ x : Set.{u}, @acc Class.{u} (∈) ↑x, @@ -1034,6 +1037,15 @@ end @[simp] theorem sUnion_empty : ⋃₀ (∅ : Class.{u}) = (∅ : Class.{u}) := by { ext, simp } +/-- An induction principle for sets. If every subset of a class is a member, then the class is + universal. -/ +theorem eq_univ_of_powerset_subset {A : Class} (hA : powerset A ⊆ A) : A = univ := +eq_univ_of_forall begin + by_contra' hnA, + exact well_founded.min_mem Set.mem_wf _ hnA (hA $ λ x hx, not_not.1 $ + λ hB, well_founded.not_lt_min Set.mem_wf _ hnA hB $ (mem_hom_right _ _).1 hx) +end + /-- The definite description operator, which is `{x}` if `{y | A y} = {x}` and `∅` otherwise. -/ def iota (A : Class) : Class := ⋃₀ {x | ∀ y, A y ↔ y = x} From a51ce54990930d581a2cb15d5299b906b5e4fb4c Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 28 Feb 2023 10:55:23 +0000 Subject: [PATCH 0557/1291] chore(data/set/semiring): add an `idem_semiring` instance (#18449) This typeclass is reasonably new, so some obvious instances are missing. Also adds some basic API lemmas that were missing about the casting functions `set.up` and `set_semiring.down`. forward-ported as https://github.com/leanprover-community/mathlib4/pull/2518 --- src/data/set/semiring.lean | 46 +++++++++++++++++++++++++++++++++----- 1 file changed, 41 insertions(+), 5 deletions(-) diff --git a/src/data/set/semiring.lean b/src/data/set/semiring.lean index 22ef765c68d2c..060fc55731325 100644 --- a/src/data/set/semiring.lean +++ b/src/data/set/semiring.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 -/ +import algebra.order.kleene import data.set.pointwise.smul /-! @@ -51,6 +52,18 @@ instance : add_comm_monoid (set_semiring α) := add_zero := union_empty, add_comm := union_comm } +lemma zero_def : (0 : set_semiring α) = set.up ∅ := rfl + +@[simp] lemma down_zero : (0 : set_semiring α).down = ∅ := rfl + +@[simp] lemma _root_.set.up_empty : (∅ : set α).up = 0 := rfl + +lemma add_def (s t : set_semiring α) : s + t = (s.down ∪ t.down).up := rfl + +@[simp] lemma down_add (s t : set_semiring α) : (s + t).down = s.down ∪ t.down := rfl + +@[simp] lemma _root_.set.up_union (s t : set α) : (s ∪ t).up = s.up + t.up := rfl + /- Since addition on `set_semiring` is commutative (it is set union), there is no need to also have the instance `covariant_class (set_semiring α) (set_semiring α) (swap (+)) (≤)`. -/ instance covariant_class_add : covariant_class (set_semiring α) (set_semiring α) (+) (≤) := @@ -60,13 +73,20 @@ section has_mul variables [has_mul α] instance : non_unital_non_assoc_semiring (set_semiring α) := -{ mul := λ s t, (image2 (*) s.down t.down).up, +{ -- reducibility linter complains if we use `(s.down * t.down).up` + mul := λ s t, (image2 (*) s.down t.down).up, zero_mul := λ s, empty_mul, mul_zero := λ s, mul_empty, left_distrib := λ _ _ _, mul_union, right_distrib := λ _ _ _, union_mul, ..set_semiring.add_comm_monoid } +lemma mul_def (s t : set_semiring α) : s * t = (s.down * t.down).up := rfl + +@[simp] lemma down_mul (s t : set_semiring α) : (s + t).down = s.down ∪ t.down := rfl + +@[simp] lemma _root_.set.up_mul (s t : set α) : (s * t).up = s.up * t.up := rfl + instance : no_zero_divisors (set_semiring α) := ⟨λ a b ab, a.eq_empty_or_nonempty.imp_right $ λ ha, b.eq_empty_or_nonempty.resolve_right $ λ hb, nonempty.ne_empty ⟨_, mul_mem_mul ha.some_mem hb.some_mem⟩ ab⟩ @@ -80,6 +100,19 @@ instance covariant_class_mul_right : end has_mul +section has_one +variables [has_one α] + +instance : has_one (set_semiring α) := { one := set.up 1 } + +lemma one_def : (1 : set_semiring α) = set.up 1 := rfl + +@[simp] lemma down_one : (1 : set_semiring α).down = 1 := rfl + +@[simp] lemma _root_.set.up_one : (1 : set α).up = 1 := rfl + +end has_one + instance [mul_one_class α] : non_assoc_semiring (set_semiring α) := { one := 1, mul := (*), @@ -88,18 +121,21 @@ instance [mul_one_class α] : non_assoc_semiring (set_semiring α) := instance [semigroup α] : non_unital_semiring (set_semiring α) := { ..set_semiring.non_unital_non_assoc_semiring, ..set.semigroup } -instance [monoid α] : semiring (set_semiring α) := -{ ..set_semiring.non_assoc_semiring, ..set_semiring.non_unital_semiring } +instance [monoid α] : idem_semiring (set_semiring α) := +{ ..set_semiring.non_assoc_semiring, ..set_semiring.non_unital_semiring, + ..set.complete_boolean_algebra } instance [comm_semigroup α] : non_unital_comm_semiring (set_semiring α) := { ..set_semiring.non_unital_semiring, ..set.comm_semigroup } +instance [comm_monoid α] : idem_comm_semiring (set_semiring α) := +{ ..set_semiring.idem_semiring, ..set.comm_monoid } + instance [comm_monoid α] : canonically_ordered_comm_semiring (set_semiring α) := { add_le_add_left := λ a b, add_le_add_left, exists_add_of_le := λ a b ab, ⟨b, (union_eq_right_iff_subset.2 ab).symm⟩, le_self_add := subset_union_left, - ..set_semiring.semiring, ..set.comm_monoid, ..set_semiring.partial_order _, - ..set_semiring.order_bot _, ..set_semiring.no_zero_divisors } + ..set_semiring.idem_semiring, ..set.comm_monoid, ..set_semiring.no_zero_divisors } /-- The image of a set under a multiplicative homomorphism is a ring homomorphism with respect to the pointwise operations on sets. -/ From 536c256e5de12c1dc0352b6b60b44f3c6c5ef340 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 28 Feb 2023 14:47:28 +0000 Subject: [PATCH 0558/1291] feat(algebra/dual_quaternion): two equivalent ways to construct the dual quaternions (#18383) We show that the dual numbers with quaternion coefficients are isomorphic as an algebra to the quaternions with dual number coefficients. The result is trivial, but being able to state it turned out to require generalizations of existing typeclass instances, which are handled in prework PRs. This is very similar to #18386. --- src/algebra/dual_quaternion.lean | 94 ++++++++++++++++++++++++++++++++ 1 file changed, 94 insertions(+) create mode 100644 src/algebra/dual_quaternion.lean diff --git a/src/algebra/dual_quaternion.lean b/src/algebra/dual_quaternion.lean new file mode 100644 index 0000000000000..863840d77d885 --- /dev/null +++ b/src/algebra/dual_quaternion.lean @@ -0,0 +1,94 @@ +/- +Copyright (c) 2023 Eric Wieser. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Eric Wieser +-/ +import algebra.dual_number +import algebra.quaternion + +/-! +# Dual quaternions + +Similar to the way that rotations in 3D space can be represented by quaternions of unit length, +rigid motions in 3D space can be represented by dual quaternions of unit length. + +## Main results + +* `quaternion.dual_number_equiv`: quaternions over dual numbers or dual + numbers over quaternions are equivalent constructions. + +## References + +* +-/ + +variables {R : Type*} [comm_ring R] + +namespace quaternion + +/-- The dual quaternions can be equivalently represented as a quaternion with dual coefficients, +or as a dual number with quaternion coefficients. + +See also `matrix.dual_number_equiv` for a similar result. -/ +def dual_number_equiv : + quaternion (dual_number R) ≃ₐ[R] dual_number (quaternion R) := +{ to_fun := λ q, + (⟨q.re.fst, q.im_i.fst, q.im_j.fst, q.im_k.fst⟩, + ⟨q.re.snd, q.im_i.snd, q.im_j.snd, q.im_k.snd⟩), + inv_fun := λ d, + ⟨(d.fst.re, d.snd.re), (d.fst.im_i, d.snd.im_i), + (d.fst.im_j, d.snd.im_j), (d.fst.im_k, d.snd.im_k)⟩, + left_inv := λ ⟨⟨r, rε⟩, ⟨i, iε⟩, ⟨j, jε⟩, ⟨k, kε⟩⟩, rfl, + right_inv := λ ⟨⟨r, i, j, k⟩, ⟨rε, iε, jε, kε⟩⟩, rfl, + map_mul' := begin + rintros ⟨⟨xr, xrε⟩, ⟨xi, xiε⟩, ⟨xj, xjε⟩, ⟨xk, xkε⟩⟩, + rintros ⟨⟨yr, yrε⟩, ⟨yi, yiε⟩, ⟨yj, yjε⟩, ⟨yk, ykε⟩⟩, + ext : 1, + { refl }, + { dsimp, + congr' 1; ring }, + end, + map_add' := begin + rintros ⟨⟨xr, xrε⟩, ⟨xi, xiε⟩, ⟨xj, xjε⟩, ⟨xk, xkε⟩⟩, + rintros ⟨⟨yr, yrε⟩, ⟨yi, yiε⟩, ⟨yj, yjε⟩, ⟨yk, ykε⟩⟩, + refl + end, + commutes' := λ r, rfl } + +/-! Lemmas characterizing `quaternion.dual_number_equiv`. -/ + +-- `simps` can't work on `dual_number` because it's not a structure +@[simp] lemma re_fst_dual_number_equiv (q : quaternion (dual_number R)) : + (dual_number_equiv q).fst.re = q.re.fst := rfl +@[simp] lemma im_i_fst_dual_number_equiv (q : quaternion (dual_number R)) : + (dual_number_equiv q).fst.im_i = q.im_i.fst := rfl +@[simp] lemma im_j_fst_dual_number_equiv (q : quaternion (dual_number R)) : + (dual_number_equiv q).fst.im_j = q.im_j.fst := rfl +@[simp] lemma im_k_fst_dual_number_equiv (q : quaternion (dual_number R)) : + (dual_number_equiv q).fst.im_k = q.im_k.fst := rfl +@[simp] lemma re_snd_dual_number_equiv (q : quaternion (dual_number R)) : + (dual_number_equiv q).snd.re = q.re.snd := rfl +@[simp] lemma im_i_snd_dual_number_equiv (q : quaternion (dual_number R)) : + (dual_number_equiv q).snd.im_i = q.im_i.snd := rfl +@[simp] lemma im_j_snd_dual_number_equiv (q : quaternion (dual_number R)) : + (dual_number_equiv q).snd.im_j = q.im_j.snd := rfl +@[simp] lemma im_k_snd_dual_number_equiv (q : quaternion (dual_number R)) : + (dual_number_equiv q).snd.im_k = q.im_k.snd := rfl +@[simp] lemma fst_re_dual_number_equiv_symm (d : dual_number (quaternion R)) : + (dual_number_equiv.symm d).re.fst = d.fst.re := rfl +@[simp] lemma fst_im_i_dual_number_equiv_symm (d : dual_number (quaternion R)) : + (dual_number_equiv.symm d).im_i.fst = d.fst.im_i := rfl +@[simp] lemma fst_im_j_dual_number_equiv_symm (d : dual_number (quaternion R)) : + (dual_number_equiv.symm d).im_j.fst = d.fst.im_j := rfl +@[simp] lemma fst_im_k_dual_number_equiv_symm (d : dual_number (quaternion R)) : + (dual_number_equiv.symm d).im_k.fst = d.fst.im_k := rfl +@[simp] lemma snd_re_dual_number_equiv_symm (d : dual_number (quaternion R)) : + (dual_number_equiv.symm d).re.snd = d.snd.re := rfl +@[simp] lemma snd_im_i_dual_number_equiv_symm (d : dual_number (quaternion R)) : + (dual_number_equiv.symm d).im_i.snd = d.snd.im_i := rfl +@[simp] lemma snd_im_j_dual_number_equiv_symm (d : dual_number (quaternion R)) : + (dual_number_equiv.symm d).im_j.snd = d.snd.im_j := rfl +@[simp] lemma snd_im_k_dual_number_equiv_symm (d : dual_number (quaternion R)) : + (dual_number_equiv.symm d).im_k.snd = d.snd.im_k := rfl + +end quaternion From da083dad7f7b82ebc6194306e9abfbff35fda260 Mon Sep 17 00:00:00 2001 From: Joanna Choules Date: Tue, 28 Feb 2023 15:56:07 +0000 Subject: [PATCH 0559/1291] feat(combinatorics/simple_graph/hom): add basic definitions and lemmas for finite subgraphs (#16382) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We define `finsubgraph` and its corresponding vertex- and edge-based singletons, and prove some basic ordering lemmas on the singletons. This is primarily preparatory work for proving the De Bruijn-Erdős theorem (generalised to arbitrary finite codomain graphs). H -----> M --> H --> E +``` +within the set `range I ⊆ E` at `I (i x) : E`. +This defines a smooth vector bundle `tangent_bundle` with fibers `tangent_space`. + +## Main definitions + +* `tangent_space I M x` is the fiber of the tangent bundle at `x : M`, which is defined to be `E`. + +* `tangent_bundle I M` is the total space of `tangent_space I M`, proven to be a smooth vector + bundle. +-/ + +open bundle set smooth_manifold_with_corners local_homeomorph +open_locale manifold topology + +noncomputable theory + +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] +{E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] +{H : Type*} [topological_space H] {I : model_with_corners 𝕜 E H} +{M : Type*} [topological_space M] [charted_space H M] [smooth_manifold_with_corners I M] +{F : Type*} [normed_add_comm_group F] [normed_space 𝕜 F] +variables (I) + +namespace hidden /- we temporarily put everything in a namespace to not have name clashes with +the existing `tangent_bundle_core`. -/ + +/-- Auxiliary lemma for tangent spaces: the derivative of a coordinate change between two charts is + smooth on its source. -/ +lemma cont_diff_on_fderiv_coord_change (i j : atlas H M) : + cont_diff_on 𝕜 ∞ (fderiv_within 𝕜 (j.1.extend I ∘ (i.1.extend I).symm) (range I)) + ((i.1.extend I).symm ≫ j.1.extend I).source := +begin + have h : ((i.1.extend I).symm ≫ j.1.extend I).source ⊆ range I, + { rw [i.1.extend_coord_change_source], apply image_subset_range }, + intros x hx, + refine (cont_diff_within_at.fderiv_within_right _ I.unique_diff le_top $ h hx).mono h, + refine (local_homeomorph.cont_diff_on_extend_coord_change I (subset_maximal_atlas I j.2) + (subset_maximal_atlas I i.2) x hx).mono_of_mem _, + exact i.1.extend_coord_change_source_mem_nhds_within j.1 I hx +end + +variables (M) +open smooth_manifold_with_corners + +/-- +Let `M` be a smooth manifold with corners with model `I` on `(E, H)`. +Then `vector_bundle_core I M` is the vector bundle core for the tangent bundle over `M`. +It is indexed by the atlas of `M`, with fiber `E` and its change of coordinates from the chart `i` +to the chart `j` at point `x : M` is the derivative of the composite +``` + I.symm i.symm j I +E -----> H -----> M --> H --> E +``` +within the set `range I ⊆ E` at `I (i x) : E`. -/ +@[simps] def tangent_bundle_core : vector_bundle_core 𝕜 M E (atlas H M) := +{ base_set := λ i, i.1.source, + is_open_base_set := λ i, i.1.open_source, + index_at := achart H, + mem_base_set_at := mem_chart_source H, + coord_change := λ i j x, fderiv_within 𝕜 (j.1.extend I ∘ (i.1.extend I).symm) (range I) + (i.1.extend I x), + coord_change_self := λ i x hx v, begin + rw [filter.eventually_eq.fderiv_within_eq, fderiv_within_id', continuous_linear_map.id_apply], + { exact I.unique_diff_at_image }, + { exact I.unique_diff_at_image }, + { filter_upwards [i.1.extend_target_mem_nhds_within I hx] with y hy, + exact (i.1.extend I).right_inv hy }, + { simp_rw [function.comp_apply, i.1.extend_left_inv I hx] } + end, + continuous_on_coord_change := λ i j, begin + refine (cont_diff_on_fderiv_coord_change I i j).continuous_on.comp + ((i.1.continuous_on_extend I).mono _) _, + { rw [i.1.extend_source], exact inter_subset_left _ _ }, + simp_rw [← i.1.extend_image_source_inter, maps_to_image] + end, + coord_change_comp := begin + rintro i j k x ⟨⟨hxi, hxj⟩, hxk⟩ v, + rw [fderiv_within_fderiv_within, filter.eventually_eq.fderiv_within_eq], + { exact I.unique_diff_at_image }, + { have := i.1.extend_preimage_mem_nhds I hxi (j.1.extend_source_mem_nhds I hxj), + filter_upwards [nhds_within_le_nhds this] with y hy, + simp_rw [function.comp_apply, (j.1.extend I).left_inv hy] }, + { simp_rw [function.comp_apply, i.1.extend_left_inv I hxi, j.1.extend_left_inv I hxj] }, + { exact (cont_diff_within_at_extend_coord_change' I (subset_maximal_atlas I k.2) + (subset_maximal_atlas I j.2) hxk hxj).differentiable_within_at le_top }, + { exact (cont_diff_within_at_extend_coord_change' I (subset_maximal_atlas I j.2) + (subset_maximal_atlas I i.2) hxj hxi).differentiable_within_at le_top }, + { intros x hx, exact mem_range_self _ }, + { exact I.unique_diff_at_image }, + { rw [function.comp_apply, i.1.extend_left_inv I hxi] } + end } + +variables {M} + +lemma tangent_bundle_core_coord_change_achart (x x' z : M) : + (tangent_bundle_core I M).coord_change (achart H x) (achart H x') z = + fderiv_within 𝕜 (ext_chart_at I x' ∘ (ext_chart_at I x).symm) (range I) (ext_chart_at I x z) := +rfl + +include I + +/-- The tangent space at a point of the manifold `M`. It is just `E`. We could use instead +`(tangent_bundle_core I M).to_topological_vector_bundle_core.fiber x`, but we use `E` to help the +kernel. +-/ +@[nolint unused_arguments, derive [topological_space, add_comm_group, topological_add_group]] +def tangent_space (x : M) : Type* := E + +omit I +variable (M) + +/-- The tangent bundle to a smooth manifold, as a Sigma type. Defined in terms of +`bundle.total_space` to be able to put a suitable topology on it. -/ +@[nolint has_nonempty_instance, reducible] -- is empty if the base manifold is empty +def tangent_bundle := bundle.total_space (tangent_space I : M → Type*) + +local notation `TM` := tangent_bundle I M + + +section tangent_bundle_instances + +/- In general, the definition of tangent_space is not reducible, so that type class inference +does not pick wrong instances. In this section, we record the right instances for +them, noting in particular that the tangent bundle is a smooth manifold. -/ + +section +local attribute [reducible] tangent_space + +variables {M} (x : M) + +instance : module 𝕜 (tangent_space I x) := by apply_instance +instance : inhabited (tangent_space I x) := ⟨0⟩ + +end + +instance : topological_space TM := +(tangent_bundle_core I M).to_fiber_bundle_core.to_topological_space + +instance : fiber_bundle E (tangent_space I : M → Type*) := +(tangent_bundle_core I M).to_fiber_bundle_core.fiber_bundle + +instance : vector_bundle 𝕜 E (tangent_space I : M → Type*) := +(tangent_bundle_core I M).vector_bundle + +lemma tangent_space_chart_at (p : TM) : + chart_at (model_prod H E) p = + ((tangent_bundle_core I M).to_fiber_bundle_core.local_triv (achart H p.1)) + .to_local_homeomorph ≫ₕ (chart_at H p.1).prod (local_homeomorph.refl E) := +rfl + +lemma tangent_space_chart_at_to_local_equiv (p : TM) : + (chart_at (model_prod H E) p).to_local_equiv = + (tangent_bundle_core I M).to_fiber_bundle_core.local_triv_as_local_equiv (achart H p.1) ≫ + (chart_at H p.1).to_local_equiv.prod (local_equiv.refl E) := +rfl + +instance tangent_bundle_core.is_smooth : (tangent_bundle_core I M).is_smooth I := +begin + refine ⟨λ i j, _⟩, + rw [smooth_on, cont_mdiff_on_iff_source_of_mem_maximal_atlas + (subset_maximal_atlas I i.2), cont_mdiff_on_iff_cont_diff_on], + refine ((cont_diff_on_fderiv_coord_change I i j).congr $ λ x hx, _).mono _, + { rw [local_equiv.trans_source'] at hx, + simp_rw [function.comp_apply, tangent_bundle_core_coord_change, + (i.1.extend I).right_inv hx.1] }, + { exact (i.1.extend_image_source_inter j.1 I).subset }, + { apply inter_subset_left } +end + +instance tangent_bundle.smooth_vector_bundle : + smooth_vector_bundle E (tangent_space I : M → Type*) I := +(tangent_bundle_core I M).smooth_vector_bundle _ + +end tangent_bundle_instances + +/-! ## The tangent bundle to the model space -/ + +/-- In the tangent bundle to the model space, the charts are just the canonical identification +between a product type and a sigma type, a.k.a. `equiv.sigma_equiv_prod`. -/ +@[simp, mfld_simps] lemma tangent_bundle_model_space_chart_at (p : tangent_bundle I H) : + (chart_at (model_prod H E) p).to_local_equiv = (equiv.sigma_equiv_prod H E).to_local_equiv := +begin + ext x : 1, + { ext, { refl }, + exact (tangent_bundle_core I H).coord_change_self (achart _ x.1) x.1 + (mem_achart_source H x.1) x.2 }, + { intros x, ext, { refl }, apply heq_of_eq, + exact (tangent_bundle_core I H).coord_change_self (achart _ x.1) x.1 + (mem_achart_source H x.1) x.2 }, + simp_rw [tangent_space_chart_at, fiber_bundle_core.local_triv, + fiber_bundle_core.local_triv_as_local_equiv, vector_bundle_core.to_fiber_bundle_core_base_set, + tangent_bundle_core_base_set], + simp only with mfld_simps, +end + +@[simp, mfld_simps] lemma tangent_bundle_model_space_coe_chart_at (p : tangent_bundle I H) : + ⇑(chart_at (model_prod H E) p) = equiv.sigma_equiv_prod H E := +by { unfold_coes, simp_rw [tangent_bundle_model_space_chart_at], refl } + +@[simp, mfld_simps] lemma tangent_bundle_model_space_coe_chart_at_symm (p : tangent_bundle I H) : + ((chart_at (model_prod H E) p).symm : model_prod H E → tangent_bundle I H) = + (equiv.sigma_equiv_prod H E).symm := +by { unfold_coes, + simp_rw [local_homeomorph.symm_to_local_equiv, tangent_bundle_model_space_chart_at], refl } + +variable (H) +/-- The canonical identification between the tangent bundle to the model space and the product, +as a homeomorphism -/ +def tangent_bundle_model_space_homeomorph : tangent_bundle I H ≃ₜ model_prod H E := +{ continuous_to_fun := + begin + let p : tangent_bundle I H := ⟨I.symm (0 : E), (0 : E)⟩, + have : continuous (chart_at (model_prod H E) p), + { rw continuous_iff_continuous_on_univ, + convert local_homeomorph.continuous_on _, + simp only [tangent_space.fiber_bundle] with mfld_simps }, + simpa only with mfld_simps using this, + end, + continuous_inv_fun := + begin + let p : tangent_bundle I H := ⟨I.symm (0 : E), (0 : E)⟩, + have : continuous (chart_at (model_prod H E) p).symm, + { rw continuous_iff_continuous_on_univ, + convert local_homeomorph.continuous_on _, + simp only with mfld_simps }, + simpa only with mfld_simps using this, + end, + .. equiv.sigma_equiv_prod H E } + +@[simp, mfld_simps] lemma tangent_bundle_model_space_homeomorph_coe : + (tangent_bundle_model_space_homeomorph H I : tangent_bundle I H → model_prod H E) + = equiv.sigma_equiv_prod H E := +rfl + +@[simp, mfld_simps] lemma tangent_bundle_model_space_homeomorph_coe_symm : + ((tangent_bundle_model_space_homeomorph H I).symm : model_prod H E → tangent_bundle I H) + = (equiv.sigma_equiv_prod H E).symm := +rfl + +end hidden From 4ac69b290818724c159de091daa3acd31da0ee6d Mon Sep 17 00:00:00 2001 From: Devon Tuma Date: Fri, 10 Mar 2023 09:55:21 +0000 Subject: [PATCH 0611/1291] feat(probability/probability_mass_function): basic lawful monad lemmas for `pmf`. (#18469) This file adds basic lemmas for monadic operations on `pmf`, mirroring the lemmas for `is_lawful_monad`. --- .../probability_mass_function/basic.lean | 16 +++++++++++++--- .../constructions.lean | 18 +++++++++++++----- .../probability_mass_function/monad.lean | 12 +++++++++--- 3 files changed, 35 insertions(+), 11 deletions(-) diff --git a/src/probability/probability_mass_function/basic.lean b/src/probability/probability_mass_function/basic.lean index 12b1c45f657ce..b468119385902 100644 --- a/src/probability/probability_mass_function/basic.lean +++ b/src/probability/probability_mass_function/basic.lean @@ -38,10 +38,13 @@ def {u} pmf (α : Type u) : Type u := { f : α → ℝ≥0∞ // has_sum f 1 } namespace pmf -instance : has_coe_to_fun (pmf α) (λ p, α → ℝ≥0∞) := ⟨λ p a, p.1 a⟩ +instance fun_like : fun_like (pmf α) α (λ p, ℝ≥0∞) := +{ coe := λ p a, p.1 a, + coe_injective' := λ p q h, subtype.eq h } -@[ext] protected lemma ext : ∀ {p q : pmf α}, (∀ a, p a = q a) → p = q -| ⟨f, hf⟩ ⟨g, hg⟩ eq := subtype.eq $ funext eq +@[ext] protected lemma ext {p q : pmf α} (h : ∀ x, p x = q x) : p = q := fun_like.ext p q h + +lemma ext_iff {p q : pmf α} : p = q ↔ ∀ x, p x = q x := fun_like.ext_iff lemma has_sum_coe_one (p : pmf α) : has_sum p 1 := p.2 @@ -53,11 +56,18 @@ lemma tsum_coe_indicator_ne_top (p : pmf α) (s : set α) : ∑' a, s.indicator ne_of_lt (lt_of_le_of_lt (tsum_le_tsum (λ a, set.indicator_apply_le (λ _, le_rfl)) ennreal.summable ennreal.summable) (lt_of_le_of_ne le_top p.tsum_coe_ne_top)) +@[simp] lemma coe_ne_zero (p : pmf α) : ⇑p ≠ 0 := +λ hp, zero_ne_one ((tsum_zero.symm.trans (tsum_congr $ + λ x, symm (congr_fun hp x))).trans p.tsum_coe) + /-- The support of a `pmf` is the set where it is nonzero. -/ def support (p : pmf α) : set α := function.support p @[simp] lemma mem_support_iff (p : pmf α) (a : α) : a ∈ p.support ↔ p a ≠ 0 := iff.rfl +@[simp] lemma support_nonempty (p : pmf α) : p.support.nonempty := +function.support_nonempty_iff.2 p.coe_ne_zero + lemma apply_eq_zero_iff (p : pmf α) (a : α) : p a = 0 ↔ a ∉ p.support := by rw [mem_support_iff, not_not] diff --git a/src/probability/probability_mass_function/constructions.lean b/src/probability/probability_mass_function/constructions.lean index 01c74b53850eb..ae7df339f0f93 100644 --- a/src/probability/probability_mass_function/constructions.lean +++ b/src/probability/probability_mass_function/constructions.lean @@ -47,13 +47,21 @@ lemma mem_support_map_iff : b ∈ (map f p).support ↔ ∃ a ∈ p.support, f a lemma bind_pure_comp : bind p (pure ∘ f) = map f p := rfl -lemma map_id : map id p = p := by simp [map] +lemma map_id : map id p = p := bind_pure _ -lemma map_comp (g : β → γ) : (p.map f).map g = p.map (g ∘ f) := -by simp [map] +lemma map_comp (g : β → γ) : (p.map f).map g = p.map (g ∘ f) := by simp [map] -lemma pure_map (a : α) : (pure a).map f = pure (f a) := -by simp [map] +lemma pure_map (a : α) : (pure a).map f = pure (f a) := pure_bind _ _ + +lemma map_bind (q : α → pmf β) (f : β → γ) : + (p.bind q).map f = p.bind (λ a, (q a).map f) := bind_bind _ _ _ + +@[simp] lemma bind_map (p : pmf α) (f : α → β) (q : β → pmf γ) : + (p.map f).bind q = p.bind (q ∘ f) := +(bind_bind _ _ _).trans (congr_arg _ (funext (λ a, pure_bind _ _))) + +@[simp] lemma map_const : p.map (function.const α b) = pure b := +by simp only [map, bind_const, function.comp_const] section measure diff --git a/src/probability/probability_mass_function/monad.lean b/src/probability/probability_mass_function/monad.lean index 042c4d0da2c0e..08d2f85af463e 100644 --- a/src/probability/probability_mass_function/monad.lean +++ b/src/probability/probability_mass_function/monad.lean @@ -39,6 +39,10 @@ variables (a a' : α) lemma mem_support_pure_iff: a' ∈ (pure a).support ↔ a' = a := by simp +@[simp] lemma pure_apply_self : pure a a = 1 := if_pos rfl + +lemma pure_apply_of_ne (h : a' ≠ a) : pure a a' = 0 := if_neg h + instance [inhabited α] : inhabited (pmf α) := ⟨pure default⟩ section measure @@ -96,9 +100,11 @@ have ∀ b a', ite (a' = a) 1 0 * f a' b = ite (a' = a) (f a b) 0, from by ext b; simp [this] @[simp] lemma bind_pure : p.bind pure = p := -have ∀ a a', (p a * ite (a' = a) 1 0) = ite (a = a') (p a') 0, from - assume a a', begin split_ifs; try { subst a }; try { subst a' }; simp * at * end, -by ext b; simp [this] +pmf.ext (λ x, (bind_apply _ _ _).trans (trans (tsum_eq_single x $ + (λ y hy, by rw [pure_apply_of_ne _ _ hy.symm, mul_zero])) $ by rw [pure_apply_self, mul_one])) + +@[simp] lemma bind_const (p : pmf α) (q : pmf β) : p.bind (λ _, q) = q := +pmf.ext (λ x, by rw [bind_apply, ennreal.tsum_mul_right, tsum_coe, one_mul]) @[simp] lemma bind_bind : (p.bind f).bind g = p.bind (λ a, (f a).bind g) := pmf.ext (λ b, by simpa only [ennreal.coe_eq_coe.symm, bind_apply, ennreal.tsum_mul_left.symm, From 3180fab693e2cee3bff62675571264cb8778b212 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Fri, 10 Mar 2023 11:38:10 +0000 Subject: [PATCH 0612/1291] feat(probability/kernel/composition): composition of kernels (#17974) We define the product, map, comap and composition of s-finite kernels. --- src/probability/kernel/composition.lean | 726 ++++++++++++++++++++++++ 1 file changed, 726 insertions(+) create mode 100644 src/probability/kernel/composition.lean diff --git a/src/probability/kernel/composition.lean b/src/probability/kernel/composition.lean new file mode 100644 index 0000000000000..68b00d917913f --- /dev/null +++ b/src/probability/kernel/composition.lean @@ -0,0 +1,726 @@ +/- +Copyright (c) 2023 Rémy Degenne. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Rémy Degenne +-/ + +import probability.kernel.basic + +/-! +# Product and composition of kernels + +We define +* the composition-product `κ ⊗ₖ η` of two s-finite kernels `κ : kernel α β` and + `η : kernel (α × β) γ`, a kernel from `α` to `β × γ`. +* the map and comap of a kernel along a measurable function. +* the composition `η ∘ₖ κ` of s-finite kernels `κ : kernel α β` and `η : kernel β γ`, + a kernel from `α` to `γ`. +* the product `prod κ η` of s-finite kernels `κ : kernel α β` and `η : kernel α γ`, + a kernel from `α` to `β × γ`. + +A note on names: +The composition-product `kernel α β → kernel (α × β) γ → kernel α (β × γ)` is named composition in +[kallenberg2021] and product on the wikipedia article on transition kernels. +Most papers studying categories of kernels call composition the map we call composition. We adopt +that convention because it fits better with the use of the name `comp` elsewhere in mathlib. + +## Main definitions + +Kernels built from other kernels: +* `comp_prod (κ : kernel α β) (η : kernel (α × β) γ) : kernel α (β × γ)`: composition-product of 2 + s-finite kernels. We define a notation `κ ⊗ₖ η = comp_prod κ η`. + `∫⁻ bc, f bc ∂((κ ⊗ₖ η) a) = ∫⁻ b, ∫⁻ c, f (b, c) ∂(η (a, b)) ∂(κ a)` +* `map (κ : kernel α β) (f : β → γ) (hf : measurable f) : kernel α γ` + `∫⁻ c, g c ∂(map κ f hf a) = ∫⁻ b, g (f b) ∂(κ a)` +* `comap (κ : kernel α β) (f : γ → α) (hf : measurable f) : kernel γ β` + `∫⁻ b, g b ∂(comap κ f hf c) = ∫⁻ b, g b ∂(κ (f c))` +* `comp (η : kernel β γ) (κ : kernel α β) : kernel α γ`: composition of 2 s-finite kernels. + We define a notation `η ∘ₖ κ = comp η κ`. + `∫⁻ c, g c ∂((η ∘ₖ κ) a) = ∫⁻ b, ∫⁻ c, g c ∂(η b) ∂(κ a)` +* `prod (κ : kernel α β) (η : kernel α γ) : kernel α (β × γ)`: product of 2 s-finite kernels. + `∫⁻ bc, f bc ∂(prod κ η a) = ∫⁻ b, ∫⁻ c, f (b, c) ∂(η a) ∂(κ a)` + +## Main statements + +* `lintegral_comp_prod`, `lintegral_map`, `lintegral_comap`, `lintegral_comp`, `lintegral_prod`: + Lebesgue integral of a function against a composition-product/map/comap/composition/product of + kernels. +* Instances of the form `.` where class is one of `is_markov_kernel`, + `is_finite_kernel`, `is_s_finite_kernel` and operation is one of `comp_prod`, `map`, `comap`, + `comp`, `prod`. These instances state that the three classes are stable by the various operations. + +## Notations + +* `κ ⊗ₖ η = probability_theory.kernel.comp_prod κ η` +* `η ∘ₖ κ = probability_theory.kernel.comp η κ` + +-/ + +open measure_theory + +open_locale ennreal + +namespace probability_theory + +namespace kernel + +variables {α β ι : Type*} {mα : measurable_space α} {mβ : measurable_space β} + +include mα mβ + +section composition_product + +/-! +### Composition-Product of kernels + +We define a kernel composition-product +`comp_prod : kernel α β → kernel (α × β) γ → kernel α (β × γ)`. +-/ + +variables {γ : Type*} {mγ : measurable_space γ} {s : set (β × γ)} + +include mγ + +/-- Auxiliary function for the definition of the composition-product of two kernels. +For all `a : α`, `comp_prod_fun κ η a` is a countably additive function with value zero on the empty +set, and the composition-product of kernels is defined in `kernel.comp_prod` through +`measure.of_measurable`. -/ +noncomputable +def comp_prod_fun (κ : kernel α β) (η : kernel (α × β) γ) (a : α) (s : set (β × γ)) : ℝ≥0∞ := +∫⁻ b, η (a, b) {c | (b, c) ∈ s} ∂(κ a) + +lemma comp_prod_fun_empty (κ : kernel α β) (η : kernel (α × β) γ) (a : α) : + comp_prod_fun κ η a ∅ = 0 := +by simp only [comp_prod_fun, set.mem_empty_iff_false, set.set_of_false, measure_empty, + lintegral_const, zero_mul] + +lemma comp_prod_fun_Union (κ : kernel α β) (η : kernel (α × β) γ) [is_s_finite_kernel η] (a : α) + (f : ℕ → set (β × γ)) (hf_meas : ∀ i, measurable_set (f i)) (hf_disj : pairwise (disjoint on f)) : + comp_prod_fun κ η a (⋃ i, f i) = ∑' i, comp_prod_fun κ η a (f i) := +begin + have h_Union : (λ b, η (a, b) {c : γ | (b, c) ∈ ⋃ i, f i}) + = λ b, η (a,b) (⋃ i, {c : γ | (b, c) ∈ f i}), + { ext1 b, + congr' with c, + simp only [set.mem_Union, set.supr_eq_Union, set.mem_set_of_eq], + refl, }, + rw [comp_prod_fun, h_Union], + have h_tsum : (λ b, η (a, b) (⋃ i, {c : γ | (b, c) ∈ f i})) + = λ b, ∑' i, η (a, b) {c : γ | (b, c) ∈ f i}, + { ext1 b, + rw measure_Union, + { intros i j hij s hsi hsj c hcs, + have hbci : {(b, c)} ⊆ f i, by { rw set.singleton_subset_iff, exact hsi hcs, }, + have hbcj : {(b, c)} ⊆ f j, by { rw set.singleton_subset_iff, exact hsj hcs, }, + simpa only [set.bot_eq_empty, set.le_eq_subset, set.singleton_subset_iff, + set.mem_empty_iff_false] using hf_disj hij hbci hbcj, }, + { exact λ i, (@measurable_prod_mk_left β γ _ _ b) _ (hf_meas i), }, }, + rw [h_tsum, lintegral_tsum], + { refl, }, + { intros i, + have hm : measurable_set {p : (α × β) × γ | (p.1.2, p.2) ∈ f i}, + from measurable_fst.snd.prod_mk measurable_snd (hf_meas i), + exact ((measurable_prod_mk_mem η hm).comp measurable_prod_mk_left).ae_measurable, }, +end + +lemma comp_prod_fun_tsum_right (κ : kernel α β) (η : kernel (α × β) γ) [is_s_finite_kernel η] + (a : α) (hs : measurable_set s) : + comp_prod_fun κ η a s = ∑' n, comp_prod_fun κ (seq η n) a s := +begin + simp_rw [comp_prod_fun, (measure_sum_seq η _).symm], + have : ∫⁻ b, measure.sum (λ n, seq η n (a, b)) {c : γ | (b, c) ∈ s} ∂(κ a) + = ∫⁻ b, ∑' n, seq η n (a, b) {c : γ | (b, c) ∈ s} ∂(κ a), + { congr', + ext1 b, + rw measure.sum_apply, + exact measurable_prod_mk_left hs, }, + rw [this, lintegral_tsum (λ n : ℕ, _)], + exact ((measurable_prod_mk_mem (seq η n) ((measurable_fst.snd.prod_mk measurable_snd) hs)).comp + measurable_prod_mk_left).ae_measurable, +end + +lemma comp_prod_fun_tsum_left (κ : kernel α β) (η : kernel (α × β) γ) [is_s_finite_kernel κ] + (a : α) (s : set (β × γ)) : + comp_prod_fun κ η a s = ∑' n, comp_prod_fun (seq κ n) η a s := +by simp_rw [comp_prod_fun, (measure_sum_seq κ _).symm, lintegral_sum_measure] + +lemma comp_prod_fun_eq_tsum (κ : kernel α β) [is_s_finite_kernel κ] + (η : kernel (α × β) γ) [is_s_finite_kernel η] (a : α) (hs : measurable_set s) : + comp_prod_fun κ η a s = ∑' n m, comp_prod_fun (seq κ n) (seq η m) a s := +by simp_rw [comp_prod_fun_tsum_left κ η a s, comp_prod_fun_tsum_right _ η a hs] + +/-- Auxiliary lemma for `measurable_comp_prod_fun`. -/ +lemma measurable_comp_prod_fun_of_finite (κ : kernel α β) [is_finite_kernel κ] + (η : kernel (α × β) γ) [is_finite_kernel η] (hs : measurable_set s) : + measurable (λ a, comp_prod_fun κ η a s) := +begin + simp only [comp_prod_fun], + have h_meas : measurable (function.uncurry (λ a b, η (a, b) {c : γ | (b, c) ∈ s})), + { have : function.uncurry (λ a b, η (a, b) {c : γ | (b, c) ∈ s}) + = λ p, η p {c : γ | (p.2, c) ∈ s}, + { ext1 p, + have hp_eq_mk : p = (p.fst, p.snd) := prod.mk.eta.symm, + rw [hp_eq_mk, function.uncurry_apply_pair], }, + rw this, + exact measurable_prod_mk_mem η (measurable_fst.snd.prod_mk measurable_snd hs), }, + exact measurable_lintegral κ h_meas, +end + +lemma measurable_comp_prod_fun (κ : kernel α β) [is_s_finite_kernel κ] + (η : kernel (α × β) γ) [is_s_finite_kernel η] (hs : measurable_set s) : + measurable (λ a, comp_prod_fun κ η a s) := +begin + simp_rw comp_prod_fun_tsum_right κ η _ hs, + refine measurable.ennreal_tsum (λ n, _), + simp only [comp_prod_fun], + have h_meas : measurable (function.uncurry (λ a b, seq η n (a, b) {c : γ | (b, c) ∈ s})), + { have : function.uncurry (λ a b, seq η n (a, b) {c : γ | (b, c) ∈ s}) + = λ p, seq η n p {c : γ | (p.2, c) ∈ s}, + { ext1 p, + have hp_eq_mk : p = (p.fst, p.snd) := prod.mk.eta.symm, + rw [hp_eq_mk, function.uncurry_apply_pair], }, + rw this, + exact measurable_prod_mk_mem (seq η n) (measurable_fst.snd.prod_mk measurable_snd hs), }, + exact measurable_lintegral κ h_meas, +end + +/-- Composition-Product of kernels. It verifies +`∫⁻ bc, f bc ∂(comp_prod κ η a) = ∫⁻ b, ∫⁻ c, f (b, c) ∂(η (a, b)) ∂(κ a)` +(see `lintegral_comp_prod`). -/ +noncomputable +def comp_prod (κ : kernel α β) [is_s_finite_kernel κ] + (η : kernel (α × β) γ) [is_s_finite_kernel η] : + kernel α (β × γ) := +{ val := λ a, measure.of_measurable (λ s hs, comp_prod_fun κ η a s) (comp_prod_fun_empty κ η a) + (comp_prod_fun_Union κ η a), + property := + begin + refine measure.measurable_of_measurable_coe _ (λ s hs, _), + have : (λ a, measure.of_measurable (λ s hs, comp_prod_fun κ η a s) (comp_prod_fun_empty κ η a) + (comp_prod_fun_Union κ η a) s) = λ a, comp_prod_fun κ η a s, + { ext1 a, rwa measure.of_measurable_apply, }, + rw this, + exact measurable_comp_prod_fun κ η hs, + end, } + +localized "infix (name := kernel.comp_prod) ` ⊗ₖ `:100 := probability_theory.kernel.comp_prod" in + probability_theory + +lemma comp_prod_apply_eq_comp_prod_fun (κ : kernel α β) [is_s_finite_kernel κ] + (η : kernel (α × β) γ) [is_s_finite_kernel η] (a : α) (hs : measurable_set s) : + (κ ⊗ₖ η) a s = comp_prod_fun κ η a s := +begin + rw [comp_prod], + change measure.of_measurable (λ s hs, comp_prod_fun κ η a s) (comp_prod_fun_empty κ η a) + (comp_prod_fun_Union κ η a) s = ∫⁻ b, η (a, b) {c | (b, c) ∈ s} ∂(κ a), + rw measure.of_measurable_apply _ hs, + refl, +end + +lemma comp_prod_apply (κ : kernel α β) [is_s_finite_kernel κ] (η : kernel (α × β) γ) + [is_s_finite_kernel η] (a : α) (hs : measurable_set s) : + (κ ⊗ₖ η) a s = ∫⁻ b, η (a, b) {c | (b, c) ∈ s} ∂(κ a) := +comp_prod_apply_eq_comp_prod_fun κ η a hs + +/-- Lebesgue integral against the composition-product of two kernels. -/ +theorem lintegral_comp_prod' (κ : kernel α β) [is_s_finite_kernel κ] (η : kernel (α × β) γ) + [is_s_finite_kernel η] (a : α) {f : β → γ → ℝ≥0∞} (hf : measurable (function.uncurry f)) : + ∫⁻ bc, f bc.1 bc.2 ∂((κ ⊗ₖ η) a) = ∫⁻ b, ∫⁻ c, f b c ∂(η (a, b)) ∂(κ a) := +begin + let F : ℕ → simple_func (β × γ) ℝ≥0∞ := simple_func.eapprox (function.uncurry f), + have h : ∀ a, (⨆ n, F n a) = function.uncurry f a, + from simple_func.supr_eapprox_apply (function.uncurry f) hf, + simp only [prod.forall, function.uncurry_apply_pair] at h, + simp_rw [← h, prod.mk.eta], + have h_mono : monotone F := λ i j hij b, simple_func.monotone_eapprox (function.uncurry f) hij _, + rw lintegral_supr (λ n, (F n).measurable) h_mono, + have : ∀ b, ∫⁻ c, (⨆ n, F n (b, c)) ∂(η (a, b)) = ⨆ n, ∫⁻ c, F n (b, c) ∂(η (a, b)), + { intro a, + rw lintegral_supr, + { exact λ n, (F n).measurable.comp measurable_prod_mk_left, }, + { exact λ i j hij b, h_mono hij _, }, }, + simp_rw this, + have h_some_meas_integral : ∀ f' : simple_func (β × γ) ℝ≥0∞, + measurable (λ b, ∫⁻ c, f' (b, c) ∂(η (a, b))), + { intros f', + have : (λ b, ∫⁻ c, f' (b, c) ∂(η (a, b))) + = (λ ab, ∫⁻ c, f' (ab.2, c) ∂(η (ab))) ∘ (λ b, (a, b)), + { ext1 ab, refl, }, + rw this, + refine measurable.comp _ measurable_prod_mk_left, + exact (measurable_lintegral η + ((simple_func.measurable _).comp (measurable_fst.snd.prod_mk measurable_snd))), }, + rw lintegral_supr, + rotate, + { exact λ n, h_some_meas_integral (F n), }, + { exact λ i j hij b, lintegral_mono (λ c, h_mono hij _), }, + congr, + ext1 n, + refine simple_func.induction _ _ (F n), + { intros c s hs, + simp only [simple_func.const_zero, simple_func.coe_piecewise, simple_func.coe_const, + simple_func.coe_zero, set.piecewise_eq_indicator, lintegral_indicator_const hs], + rw [comp_prod_apply κ η _ hs, ← lintegral_const_mul c _], + swap, { exact (measurable_prod_mk_mem η ((measurable_fst.snd.prod_mk measurable_snd) hs)).comp + measurable_prod_mk_left, }, + congr, + ext1 b, + rw lintegral_indicator_const_comp measurable_prod_mk_left hs, + refl, }, + { intros f f' h_disj hf_eq hf'_eq, + simp_rw [simple_func.coe_add, pi.add_apply], + change ∫⁻ x, ((f : (β × γ) → ℝ≥0∞) x + f' x) ∂((κ ⊗ₖ η) a) + = ∫⁻ b, ∫⁻ (c : γ), f (b, c) + f' (b, c) ∂(η (a, b)) ∂(κ a), + rw [lintegral_add_left (simple_func.measurable _), hf_eq, hf'_eq, ← lintegral_add_left], + swap, { exact h_some_meas_integral f, }, + congr' with b, + rw ← lintegral_add_left ((simple_func.measurable _).comp measurable_prod_mk_left), }, +end + +/-- Lebesgue integral against the composition-product of two kernels. -/ +theorem lintegral_comp_prod (κ : kernel α β) [is_s_finite_kernel κ] (η : kernel (α × β) γ) + [is_s_finite_kernel η] (a : α) {f : β × γ → ℝ≥0∞} (hf : measurable f) : + ∫⁻ bc, f bc ∂((κ ⊗ₖ η) a) = ∫⁻ b, ∫⁻ c, f (b, c) ∂(η (a, b)) ∂(κ a) := +begin + let g := function.curry f, + change ∫⁻ bc, f bc ∂((κ ⊗ₖ η) a) = ∫⁻ b, ∫⁻ c, g b c ∂(η (a, b)) ∂(κ a), + rw ← lintegral_comp_prod', + { simp_rw [g, function.curry_apply, prod.mk.eta], }, + { simp_rw [g, function.uncurry_curry], exact hf, }, +end + +lemma comp_prod_eq_tsum_comp_prod (κ : kernel α β) [is_s_finite_kernel κ] (η : kernel (α × β) γ) + [is_s_finite_kernel η] (a : α) (hs : measurable_set s) : + (κ ⊗ₖ η) a s = ∑' (n m : ℕ), (seq κ n ⊗ₖ seq η m) a s := +by { simp_rw comp_prod_apply_eq_comp_prod_fun _ _ _ hs, exact comp_prod_fun_eq_tsum κ η a hs, } + +lemma comp_prod_eq_sum_comp_prod (κ : kernel α β) [is_s_finite_kernel κ] + (η : kernel (α × β) γ) [is_s_finite_kernel η] : + κ ⊗ₖ η = kernel.sum (λ n, kernel.sum (λ m, seq κ n ⊗ₖ seq η m)) := +by { ext a s hs : 2, simp_rw [kernel.sum_apply' _ a hs], rw comp_prod_eq_tsum_comp_prod κ η a hs, } + +lemma comp_prod_eq_sum_comp_prod_left (κ : kernel α β) [is_s_finite_kernel κ] + (η : kernel (α × β) γ) [is_s_finite_kernel η] : + κ ⊗ₖ η = kernel.sum (λ n, seq κ n ⊗ₖ η) := +begin + rw comp_prod_eq_sum_comp_prod, + congr' with n a s hs, + simp_rw [kernel.sum_apply' _ _ hs, comp_prod_apply_eq_comp_prod_fun _ _ _ hs, + comp_prod_fun_tsum_right _ η a hs], +end + +lemma comp_prod_eq_sum_comp_prod_right (κ : kernel α β) [is_s_finite_kernel κ] + (η : kernel (α × β) γ) [is_s_finite_kernel η] : + κ ⊗ₖ η = kernel.sum (λ n, κ ⊗ₖ seq η n) := +begin + rw comp_prod_eq_sum_comp_prod, + simp_rw comp_prod_eq_sum_comp_prod_left κ _, + rw kernel.sum_comm, +end + +instance is_markov_kernel.comp_prod (κ : kernel α β) [is_markov_kernel κ] + (η : kernel (α × β) γ) [is_markov_kernel η] : + is_markov_kernel (κ ⊗ₖ η) := +⟨λ a, ⟨begin + rw comp_prod_apply κ η a measurable_set.univ, + simp only [set.mem_univ, set.set_of_true, measure_univ, lintegral_one], +end⟩⟩ + +lemma comp_prod_apply_univ_le (κ : kernel α β) [is_s_finite_kernel κ] + (η : kernel (α × β) γ) [is_finite_kernel η] (a : α) : + (κ ⊗ₖ η) a set.univ ≤ (κ a set.univ) * (is_finite_kernel.bound η) := +begin + rw comp_prod_apply κ η a measurable_set.univ, + simp only [set.mem_univ, set.set_of_true], + let Cη := is_finite_kernel.bound η, + calc ∫⁻ b, η (a, b) set.univ ∂(κ a) + ≤ ∫⁻ b, Cη ∂(κ a) : lintegral_mono (λ b, measure_le_bound η (a, b) set.univ) + ... = Cη * κ a set.univ : lintegral_const Cη + ... = κ a set.univ * Cη : mul_comm _ _, +end + +instance is_finite_kernel.comp_prod (κ : kernel α β) [is_finite_kernel κ] + (η : kernel (α × β) γ) [is_finite_kernel η] : + is_finite_kernel (κ ⊗ₖ η) := +⟨⟨is_finite_kernel.bound κ * is_finite_kernel.bound η, + ennreal.mul_lt_top (is_finite_kernel.bound_ne_top κ) (is_finite_kernel.bound_ne_top η), + λ a, calc (κ ⊗ₖ η) a set.univ + ≤ (κ a set.univ) * is_finite_kernel.bound η : comp_prod_apply_univ_le κ η a +... ≤ is_finite_kernel.bound κ * is_finite_kernel.bound η : + mul_le_mul (measure_le_bound κ a set.univ) le_rfl (zero_le _) (zero_le _), ⟩⟩ + +instance is_s_finite_kernel.comp_prod (κ : kernel α β) [is_s_finite_kernel κ] + (η : kernel (α × β) γ) [is_s_finite_kernel η] : + is_s_finite_kernel (κ ⊗ₖ η) := +begin + rw comp_prod_eq_sum_comp_prod, + exact kernel.is_s_finite_kernel_sum (λ n, kernel.is_s_finite_kernel_sum infer_instance), +end + +end composition_product + +section map_comap +/-! ### map, comap -/ + +variables {γ : Type*} {mγ : measurable_space γ} {f : β → γ} {g : γ → α} + +include mγ + +/-- The pushforward of a kernel along a measurable function. +We include measurability in the assumptions instead of using junk values +to make sure that typeclass inference can infer that the `map` of a Markov kernel +is again a Markov kernel. -/ +noncomputable +def map (κ : kernel α β) (f : β → γ) (hf : measurable f) : kernel α γ := +{ val := λ a, (κ a).map f, + property := (measure.measurable_map _ hf).comp (kernel.measurable κ) } + +lemma map_apply (κ : kernel α β) (hf : measurable f) (a : α) : + map κ f hf a = (κ a).map f := rfl + +lemma map_apply' (κ : kernel α β) (hf : measurable f) (a : α) {s : set γ} (hs : measurable_set s) : + map κ f hf a s = κ a (f ⁻¹' s) := +by rw [map_apply, measure.map_apply hf hs] + +lemma lintegral_map (κ : kernel α β) (hf : measurable f) (a : α) + {g' : γ → ℝ≥0∞} (hg : measurable g') : + ∫⁻ b, g' b ∂(map κ f hf a) = ∫⁻ a, g' (f a) ∂(κ a) := +by rw [map_apply _ hf, lintegral_map hg hf] + +lemma sum_map_seq (κ : kernel α β) [is_s_finite_kernel κ] (hf : measurable f) : + kernel.sum (λ n, map (seq κ n) f hf) = map κ f hf := +begin + ext a s hs : 2, + rw [kernel.sum_apply, map_apply' κ hf a hs, measure.sum_apply _ hs, ← measure_sum_seq κ, + measure.sum_apply _ (hf hs)], + simp_rw map_apply' _ hf _ hs, +end + +instance is_markov_kernel.map (κ : kernel α β) [is_markov_kernel κ] (hf : measurable f) : + is_markov_kernel (map κ f hf) := + ⟨λ a, ⟨by rw [map_apply' κ hf a measurable_set.univ, set.preimage_univ, measure_univ]⟩⟩ + +instance is_finite_kernel.map (κ : kernel α β) [is_finite_kernel κ] (hf : measurable f) : + is_finite_kernel (map κ f hf) := +begin + refine ⟨⟨is_finite_kernel.bound κ, is_finite_kernel.bound_lt_top κ, λ a, _⟩⟩, + rw map_apply' κ hf a measurable_set.univ, + exact measure_le_bound κ a _, +end + +instance is_s_finite_kernel.map (κ : kernel α β) [is_s_finite_kernel κ] (hf : measurable f) : + is_s_finite_kernel (map κ f hf) := +⟨⟨λ n, map (seq κ n) f hf, infer_instance, (sum_map_seq κ hf).symm⟩⟩ + +/-- Pullback of a kernel, such that for each set s `comap κ g hg c s = κ (g c) s`. +We include measurability in the assumptions instead of using junk values +to make sure that typeclass inference can infer that the `comap` of a Markov kernel +is again a Markov kernel. -/ +def comap (κ : kernel α β) (g : γ → α) (hg : measurable g) : kernel γ β := +{ val := λ a, κ (g a), + property := (kernel.measurable κ).comp hg } + +lemma comap_apply (κ : kernel α β) (hg : measurable g) (c : γ) : + comap κ g hg c = κ (g c) := rfl + +lemma comap_apply' (κ : kernel α β) (hg : measurable g) (c : γ) (s : set β) : + comap κ g hg c s = κ (g c) s := rfl + +lemma lintegral_comap (κ : kernel α β) (hg : measurable g) (c : γ) (g' : β → ℝ≥0∞) : + ∫⁻ b, g' b ∂(comap κ g hg c) = ∫⁻ b, g' b ∂(κ (g c)) := rfl + +lemma sum_comap_seq (κ : kernel α β) [is_s_finite_kernel κ] (hg : measurable g) : + kernel.sum (λ n, comap (seq κ n) g hg) = comap κ g hg := +begin + ext a s hs : 2, + rw [kernel.sum_apply, comap_apply' κ hg a s, measure.sum_apply _ hs, ← measure_sum_seq κ, + measure.sum_apply _ hs], + simp_rw comap_apply' _ hg _ s, +end + +instance is_markov_kernel.comap (κ : kernel α β) [is_markov_kernel κ] (hg : measurable g) : + is_markov_kernel (comap κ g hg) := +⟨λ a, ⟨by rw [comap_apply' κ hg a set.univ, measure_univ]⟩⟩ + +instance is_finite_kernel.comap (κ : kernel α β) [is_finite_kernel κ] (hg : measurable g) : + is_finite_kernel (comap κ g hg) := +begin + refine ⟨⟨is_finite_kernel.bound κ, is_finite_kernel.bound_lt_top κ, λ a, _⟩⟩, + rw comap_apply' κ hg a set.univ, + exact measure_le_bound κ _ _, +end + +instance is_s_finite_kernel.comap (κ : kernel α β) [is_s_finite_kernel κ] (hg : measurable g) : + is_s_finite_kernel (comap κ g hg) := +⟨⟨λ n, comap (seq κ n) g hg, infer_instance, (sum_comap_seq κ hg).symm⟩⟩ + +end map_comap + +open_locale probability_theory + +section fst_snd + +/-- Define a `kernel (γ × α) β` from a `kernel α β` by taking the comap of the projection. -/ +def prod_mk_left (κ : kernel α β) (γ : Type*) [measurable_space γ] : kernel (γ × α) β := +comap κ prod.snd measurable_snd + +variables {γ : Type*} {mγ : measurable_space γ} {f : β → γ} {g : γ → α} + +include mγ + +lemma prod_mk_left_apply (κ : kernel α β) (ca : γ × α) : + prod_mk_left κ γ ca = κ ca.snd := rfl + +lemma prod_mk_left_apply' (κ : kernel α β) (ca : γ × α) (s : set β) : + prod_mk_left κ γ ca s = κ ca.snd s := rfl + +lemma lintegral_prod_mk_left (κ : kernel α β) (ca : γ × α) (g : β → ℝ≥0∞) : + ∫⁻ b, g b ∂(prod_mk_left κ γ ca) = ∫⁻ b, g b ∂(κ ca.snd) := rfl + +instance is_markov_kernel.prod_mk_left (κ : kernel α β) [is_markov_kernel κ] : + is_markov_kernel (prod_mk_left κ γ) := +by { rw prod_mk_left, apply_instance, } + +instance is_finite_kernel.prod_mk_left (κ : kernel α β) [is_finite_kernel κ] : + is_finite_kernel (prod_mk_left κ γ) := +by { rw prod_mk_left, apply_instance, } + +instance is_s_finite_kernel.prod_mk_left (κ : kernel α β) [is_s_finite_kernel κ] : + is_s_finite_kernel (prod_mk_left κ γ) := +by { rw prod_mk_left, apply_instance, } + +/-- Define a `kernel (β × α) γ` from a `kernel (α × β) γ` by taking the comap of `prod.swap`. -/ +def swap_left (κ : kernel (α × β) γ) : kernel (β × α) γ := +comap κ prod.swap measurable_swap + +lemma swap_left_apply (κ : kernel (α × β) γ) (a : β × α) : + swap_left κ a = (κ a.swap) := rfl + +lemma swap_left_apply' (κ : kernel (α × β) γ) (a : β × α) (s : set γ) : + swap_left κ a s = κ a.swap s := rfl + +lemma lintegral_swap_left (κ : kernel (α × β) γ) (a : β × α) (g : γ → ℝ≥0∞) : + ∫⁻ c, g c ∂(swap_left κ a) = ∫⁻ c, g c ∂(κ a.swap) := +by { rw [swap_left, lintegral_comap _ measurable_swap a], } + +instance is_markov_kernel.swap_left (κ : kernel (α × β) γ) [is_markov_kernel κ] : + is_markov_kernel (swap_left κ) := +by { rw swap_left, apply_instance, } + +instance is_finite_kernel.swap_left (κ : kernel (α × β) γ) [is_finite_kernel κ] : + is_finite_kernel (swap_left κ) := +by { rw swap_left, apply_instance, } + +instance is_s_finite_kernel.swap_left (κ : kernel (α × β) γ) [is_s_finite_kernel κ] : + is_s_finite_kernel (swap_left κ) := +by { rw swap_left, apply_instance, } + +/-- Define a `kernel α (γ × β)` from a `kernel α (β × γ)` by taking the map of `prod.swap`. -/ +noncomputable +def swap_right (κ : kernel α (β × γ)) : kernel α (γ × β) := +map κ prod.swap measurable_swap + +lemma swap_right_apply (κ : kernel α (β × γ)) (a : α) : + swap_right κ a = (κ a).map prod.swap := rfl + +lemma swap_right_apply' (κ : kernel α (β × γ)) (a : α) {s : set (γ × β)} (hs : measurable_set s) : + swap_right κ a s = κ a {p | p.swap ∈ s} := +by { rw [swap_right_apply, measure.map_apply measurable_swap hs], refl, } + +lemma lintegral_swap_right (κ : kernel α (β × γ)) (a : α) {g : γ × β → ℝ≥0∞} (hg : measurable g) : + ∫⁻ c, g c ∂(swap_right κ a) = ∫⁻ (bc : β × γ), g bc.swap ∂(κ a) := +by rw [swap_right, lintegral_map _ measurable_swap a hg] + +instance is_markov_kernel.swap_right (κ : kernel α (β × γ)) [is_markov_kernel κ] : + is_markov_kernel (swap_right κ) := +by { rw swap_right, apply_instance, } + +instance is_finite_kernel.swap_right (κ : kernel α (β × γ)) [is_finite_kernel κ] : + is_finite_kernel (swap_right κ) := +by { rw swap_right, apply_instance, } + +instance is_s_finite_kernel.swap_right (κ : kernel α (β × γ)) [is_s_finite_kernel κ] : + is_s_finite_kernel (swap_right κ) := +by { rw swap_right, apply_instance, } + +/-- Define a `kernel α β` from a `kernel α (β × γ)` by taking the map of the first projection. -/ +noncomputable +def fst (κ : kernel α (β × γ)) : kernel α β := +map κ prod.fst measurable_fst + +lemma fst_apply (κ : kernel α (β × γ)) (a : α) : + fst κ a = (κ a).map prod.fst := rfl + +lemma fst_apply' (κ : kernel α (β × γ)) (a : α) {s : set β} (hs : measurable_set s) : + fst κ a s = κ a {p | p.1 ∈ s} := +by { rw [fst_apply, measure.map_apply measurable_fst hs], refl, } + +lemma lintegral_fst (κ : kernel α (β × γ)) (a : α) {g : β → ℝ≥0∞} (hg : measurable g) : + ∫⁻ c, g c ∂(fst κ a) = ∫⁻ (bc : β × γ), g bc.fst ∂(κ a) := +by rw [fst, lintegral_map _ measurable_fst a hg] + +instance is_markov_kernel.fst (κ : kernel α (β × γ)) [is_markov_kernel κ] : + is_markov_kernel (fst κ) := +by { rw fst, apply_instance, } + +instance is_finite_kernel.fst (κ : kernel α (β × γ)) [is_finite_kernel κ] : + is_finite_kernel (fst κ) := +by { rw fst, apply_instance, } + +instance is_s_finite_kernel.fst (κ : kernel α (β × γ)) [is_s_finite_kernel κ] : + is_s_finite_kernel (fst κ) := +by { rw fst, apply_instance, } + +/-- Define a `kernel α γ` from a `kernel α (β × γ)` by taking the map of the second projection. -/ +noncomputable +def snd (κ : kernel α (β × γ)) : kernel α γ := +map κ prod.snd measurable_snd + +lemma snd_apply (κ : kernel α (β × γ)) (a : α) : + snd κ a = (κ a).map prod.snd := rfl + +lemma snd_apply' (κ : kernel α (β × γ)) (a : α) {s : set γ} (hs : measurable_set s) : + snd κ a s = κ a {p | p.2 ∈ s} := +by { rw [snd_apply, measure.map_apply measurable_snd hs], refl, } + +lemma lintegral_snd (κ : kernel α (β × γ)) (a : α) {g : γ → ℝ≥0∞} (hg : measurable g) : + ∫⁻ c, g c ∂(snd κ a) = ∫⁻ (bc : β × γ), g bc.snd ∂(κ a) := +by rw [snd, lintegral_map _ measurable_snd a hg] + +instance is_markov_kernel.snd (κ : kernel α (β × γ)) [is_markov_kernel κ] : + is_markov_kernel (snd κ) := +by { rw snd, apply_instance, } + +instance is_finite_kernel.snd (κ : kernel α (β × γ)) [is_finite_kernel κ] : + is_finite_kernel (snd κ) := +by { rw snd, apply_instance, } + +instance is_s_finite_kernel.snd (κ : kernel α (β × γ)) [is_s_finite_kernel κ] : + is_s_finite_kernel (snd κ) := +by { rw snd, apply_instance, } + +end fst_snd + +section comp +/-! ### Composition of two kernels -/ + +variables {γ : Type*} {mγ : measurable_space γ} {f : β → γ} {g : γ → α} + +include mγ + +/-- Composition of two s-finite kernels. -/ +noncomputable +def comp (η : kernel β γ) [is_s_finite_kernel η] (κ : kernel α β) [is_s_finite_kernel κ] : + kernel α γ := +snd (κ ⊗ₖ prod_mk_left η α) + +localized "infix (name := kernel.comp) ` ∘ₖ `:100 := probability_theory.kernel.comp" in + probability_theory + +lemma comp_apply (η : kernel β γ) [is_s_finite_kernel η] (κ : kernel α β) [is_s_finite_kernel κ] + (a : α) {s : set γ} (hs : measurable_set s) : + (η ∘ₖ κ) a s = ∫⁻ b, η b s ∂(κ a) := +begin + rw [comp, snd_apply' _ _ hs, comp_prod_apply], + swap, { exact measurable_snd hs, }, + simp only [set.mem_set_of_eq, set.set_of_mem_eq, prod_mk_left_apply' _ _ s], +end + +lemma lintegral_comp (η : kernel β γ) [is_s_finite_kernel η] (κ : kernel α β) [is_s_finite_kernel κ] + (a : α) {g : γ → ℝ≥0∞} (hg : measurable g) : + ∫⁻ c, g c ∂((η ∘ₖ κ) a) = ∫⁻ b, ∫⁻ c, g c ∂(η b) ∂(κ a) := +begin + rw [comp, lintegral_snd _ _ hg], + change ∫⁻ bc, (λ a b, g b) bc.fst bc.snd ∂((κ ⊗ₖ prod_mk_left η α) a) + = ∫⁻ b, ∫⁻ c, g c ∂(η b) ∂(κ a), + exact lintegral_comp_prod _ _ _ (hg.comp measurable_snd), +end + +instance is_markov_kernel.comp (η : kernel β γ) [is_markov_kernel η] + (κ : kernel α β) [is_markov_kernel κ] : + is_markov_kernel (η ∘ₖ κ) := +by { rw comp, apply_instance, } + +instance is_finite_kernel.comp (η : kernel β γ) [is_finite_kernel η] + (κ : kernel α β) [is_finite_kernel κ] : + is_finite_kernel (η ∘ₖ κ) := +by { rw comp, apply_instance, } + +instance is_s_finite_kernel.comp (η : kernel β γ) [is_s_finite_kernel η] + (κ : kernel α β) [is_s_finite_kernel κ] : + is_s_finite_kernel (η ∘ₖ κ) := +by { rw comp, apply_instance, } + +/-- Composition of kernels is associative. -/ +lemma comp_assoc {δ : Type*} {mδ : measurable_space δ} (ξ : kernel γ δ) [is_s_finite_kernel ξ] + (η : kernel β γ) [is_s_finite_kernel η] (κ : kernel α β) [is_s_finite_kernel κ] : + (ξ ∘ₖ η ∘ₖ κ) = ξ ∘ₖ (η ∘ₖ κ) := +begin + refine ext_fun (λ a f hf, _), + simp_rw [lintegral_comp _ _ _ hf, lintegral_comp _ _ _ (measurable_lintegral' ξ hf)], +end + +lemma deterministic_comp_eq_map (hf : measurable f) (κ : kernel α β) [is_s_finite_kernel κ] : + (deterministic hf ∘ₖ κ) = map κ f hf := +begin + ext a s hs : 2, + simp_rw [map_apply' _ _ _ hs, comp_apply _ _ _ hs, deterministic_apply' hf _ hs, + lintegral_indicator_const_comp hf hs, one_mul], +end + +lemma comp_deterministic_eq_comap (κ : kernel α β) [is_s_finite_kernel κ] (hg : measurable g) : + (κ ∘ₖ deterministic hg) = comap κ g hg := +begin + ext a s hs : 2, + simp_rw [comap_apply' _ _ _ s, comp_apply _ _ _ hs, deterministic_apply hg a, + lintegral_dirac' _ (kernel.measurable_coe κ hs)], +end + +end comp + +section prod + +/-! ### Product of two kernels -/ + +variables {γ : Type*} {mγ : measurable_space γ} + +include mγ + +/-- Product of two s-finite kernels. -/ +noncomputable +def prod (κ : kernel α β) [is_s_finite_kernel κ] (η : kernel α γ) [is_s_finite_kernel η] : + kernel α (β × γ) := +κ ⊗ₖ (swap_left (prod_mk_left η β)) + +lemma prod_apply (κ : kernel α β) [is_s_finite_kernel κ] (η : kernel α γ) [is_s_finite_kernel η] + (a : α) {s : set (β × γ)} (hs : measurable_set s) : + (prod κ η) a s = ∫⁻ (b : β), (η a) {c : γ | (b, c) ∈ s} ∂(κ a) := +by simp_rw [prod, comp_prod_apply _ _ _ hs, swap_left_apply _ _, prod_mk_left_apply, + prod.swap_prod_mk] + +lemma lintegral_prod (κ : kernel α β) [is_s_finite_kernel κ] (η : kernel α γ) [is_s_finite_kernel η] + (a : α) {g : (β × γ) → ℝ≥0∞} (hg : measurable g) : + ∫⁻ c, g c ∂((prod κ η) a) = ∫⁻ b, ∫⁻ c, g (b, c) ∂(η a) ∂(κ a) := +by simp_rw [prod, lintegral_comp_prod _ _ _ hg, swap_left_apply, prod_mk_left_apply, + prod.swap_prod_mk] + +instance is_markov_kernel.prod (κ : kernel α β) [is_markov_kernel κ] + (η : kernel α γ) [is_markov_kernel η] : + is_markov_kernel (prod κ η) := +by { rw prod, apply_instance, } + +instance is_finite_kernel.prod (κ : kernel α β) [is_finite_kernel κ] + (η : kernel α γ) [is_finite_kernel η] : + is_finite_kernel (prod κ η) := +by { rw prod, apply_instance, } + +instance is_s_finite_kernel.prod (κ : kernel α β) [is_s_finite_kernel κ] + (η : kernel α γ) [is_s_finite_kernel η] : + is_s_finite_kernel (prod κ η) := +by { rw prod, apply_instance, } + +end prod + +end kernel + +end probability_theory From da420a8c6dd5bdfb85c4ced85c34388f633bc6ff Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Sat, 11 Mar 2023 09:26:40 +0000 Subject: [PATCH 0613/1291] refactor(*): reduce dependencies on ring_theory.ideal.quotient_operations (#18530) Co-authored-by: Eric Wieser --- src/data/polynomial/div.lean | 16 -- src/data/zmod/quotient.lean | 1 + .../matrix/nonsingular_inverse.lean | 6 +- src/number_theory/multiplicity.lean | 1 + src/ring_theory/adjoin_root.lean | 1 + src/ring_theory/eisenstein_criterion.lean | 1 + src/ring_theory/etale.lean | 2 +- src/ring_theory/finite_presentation.lean | 1 + src/ring_theory/jacobson_ideal.lean | 2 +- src/ring_theory/nilpotent.lean | 69 +----- src/ring_theory/noetherian.lean | 4 - src/ring_theory/polynomial/basic.lean | 221 ++++------------- src/ring_theory/polynomial/quotient.lean | 229 ++++++++++++++++++ src/ring_theory/quotient_nilpotent.lean | 74 ++++++ src/ring_theory/quotient_noetherian.lean | 15 ++ src/ring_theory/valuation/basic.lean | 84 +------ src/ring_theory/valuation/quotient.lean | 112 +++++++++ 17 files changed, 488 insertions(+), 351 deletions(-) create mode 100644 src/ring_theory/polynomial/quotient.lean create mode 100644 src/ring_theory/quotient_nilpotent.lean create mode 100644 src/ring_theory/quotient_noetherian.lean create mode 100644 src/ring_theory/valuation/quotient.lean diff --git a/src/data/polynomial/div.lean b/src/data/polynomial/div.lean index ea1343b23e460..67ab0564bd65b 100644 --- a/src/data/polynomial/div.lean +++ b/src/data/polynomial/div.lean @@ -449,22 +449,6 @@ variable {R} lemma ker_eval_ring_hom (x : R) : (eval_ring_hom x).ker = ideal.span {X - C x} := by { ext y, simpa only [ideal.mem_span_singleton, dvd_iff_is_root] } -/-- For a commutative ring $R$, evaluating a polynomial at an element $x \in R$ induces an -isomorphism of $R$-algebras $R[X] / \langle X - x \rangle \cong R$. -/ -noncomputable def quotient_span_X_sub_C_alg_equiv (x : R) : - (R[X] ⧸ ideal.span ({X - C x} : set R[X])) ≃ₐ[R] R := -(alg_equiv.restrict_scalars R $ ideal.quotient_equiv_alg_of_eq R - (by exact ker_eval_ring_hom x : ring_hom.ker (aeval x).to_ring_hom = _)).symm.trans $ - ideal.quotient_ker_alg_equiv_of_right_inverse $ λ _, eval_C - -@[simp] lemma quotient_span_X_sub_C_alg_equiv_mk (x : R) (p : R[X]) : - quotient_span_X_sub_C_alg_equiv x (ideal.quotient.mk _ p) = p.eval x := -rfl - -@[simp] lemma quotient_span_X_sub_C_alg_equiv_symm_apply (x : R) (y : R) : - (quotient_span_X_sub_C_alg_equiv x).symm y = algebra_map R _ y := -rfl - section multiplicity /-- An algorithm for deciding polynomial divisibility. The algorithm is "compute `p %ₘ q` and compare to `0`". diff --git a/src/data/zmod/quotient.lean b/src/data/zmod/quotient.lean index 0d84715b4a02f..3b4fedf4093b5 100644 --- a/src/data/zmod/quotient.lean +++ b/src/data/zmod/quotient.lean @@ -6,6 +6,7 @@ Authors: Anne Baanen import data.zmod.basic import group_theory.group_action.quotient import ring_theory.int.basic +import ring_theory.ideal.quotient_operations /-! # `zmod n` and quotient groups / rings diff --git a/src/linear_algebra/matrix/nonsingular_inverse.lean b/src/linear_algebra/matrix/nonsingular_inverse.lean index 5cadefd13c9fe..9493e921678fe 100644 --- a/src/linear_algebra/matrix/nonsingular_inverse.lean +++ b/src/linear_algebra/matrix/nonsingular_inverse.lean @@ -458,7 +458,11 @@ invertible.map (diagonal_ring_hom n α) v lemma inv_of_diagonal_eq {α} [semiring α] (v : n → α) [invertible v] [invertible (diagonal v)] : ⅟(diagonal v) = diagonal (⅟v) := -by { letI := diagonal_invertible v, convert (rfl : ⅟(diagonal v) = _) } +begin + letI := diagonal_invertible v, + haveI := invertible.subsingleton (diagonal v), + convert (rfl : ⅟(diagonal v) = _), +end /-- `v` is invertible if `diagonal v` is -/ def invertible_of_diagonal_invertible (v : n → α) [invertible (diagonal v)] : invertible v := diff --git a/src/number_theory/multiplicity.lean b/src/number_theory/multiplicity.lean index baaa7f075ecfe..82c6eed3af5f6 100644 --- a/src/number_theory/multiplicity.lean +++ b/src/number_theory/multiplicity.lean @@ -7,6 +7,7 @@ import algebra.geom_sum import data.int.parity import data.zmod.basic import number_theory.padics.padic_val +import ring_theory.ideal.quotient_operations /-! # Multiplicity in Number Theory diff --git a/src/ring_theory/adjoin_root.lean b/src/ring_theory/adjoin_root.lean index 445451c46b548..2539cba43fd04 100644 --- a/src/ring_theory/adjoin_root.lean +++ b/src/ring_theory/adjoin_root.lean @@ -11,6 +11,7 @@ import ring_theory.finite_presentation import ring_theory.finite_type import ring_theory.power_basis import ring_theory.principal_ideal_domain +import ring_theory.quotient_noetherian /-! # Adjoining roots of polynomials diff --git a/src/ring_theory/eisenstein_criterion.lean b/src/ring_theory/eisenstein_criterion.lean index 5ee839d30dc37..fe9dc886ef0b0 100644 --- a/src/ring_theory/eisenstein_criterion.lean +++ b/src/ring_theory/eisenstein_criterion.lean @@ -6,6 +6,7 @@ Authors: Chris Hughes import data.nat.cast.with_top import ring_theory.prime import ring_theory.polynomial.content +import ring_theory.ideal.quotient_operations /-! # Eisenstein's criterion diff --git a/src/ring_theory/etale.lean b/src/ring_theory/etale.lean index 01ea28e032428..f3d9a2a562c67 100644 --- a/src/ring_theory/etale.lean +++ b/src/ring_theory/etale.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Andrew Yang. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Andrew Yang -/ -import ring_theory.nilpotent +import ring_theory.quotient_nilpotent import ring_theory.derivation /-! diff --git a/src/ring_theory/finite_presentation.lean b/src/ring_theory/finite_presentation.lean index 39908602603d6..3084be7a24354 100644 --- a/src/ring_theory/finite_presentation.lean +++ b/src/ring_theory/finite_presentation.lean @@ -6,6 +6,7 @@ Authors: Johan Commelin import ring_theory.finite_type import ring_theory.mv_polynomial.tower +import ring_theory.ideal.quotient_operations /-! # Finiteness conditions in commutative algebra diff --git a/src/ring_theory/jacobson_ideal.lean b/src/ring_theory/jacobson_ideal.lean index 8055f53a97848..0e1f0e5efbbcb 100644 --- a/src/ring_theory/jacobson_ideal.lean +++ b/src/ring_theory/jacobson_ideal.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau, Devon Tuma -/ import ring_theory.ideal.quotient -import ring_theory.polynomial.basic +import ring_theory.polynomial.quotient /-! # Jacobson radical diff --git a/src/ring_theory/nilpotent.lean b/src/ring_theory/nilpotent.lean index 96fad82df57e7..775b5888b7d28 100644 --- a/src/ring_theory/nilpotent.lean +++ b/src/ring_theory/nilpotent.lean @@ -5,7 +5,7 @@ Authors: Oliver Nash -/ import data.nat.choose.sum import algebra.algebra.bilinear -import ring_theory.ideal.quotient_operations +import ring_theory.ideal.operations /-! # Nilpotent elements @@ -88,11 +88,6 @@ lemma ring_hom.ker_is_radical_iff_reduced_of_surjective {S F} [comm_semiring R] (ring_hom.ker f).is_radical ↔ is_reduced S := by simp_rw [is_reduced_iff, hf.forall, is_nilpotent, ← map_pow, ← ring_hom.mem_ker]; refl -lemma ideal.is_radical_iff_quotient_reduced [comm_ring R] (I : ideal R) : - I.is_radical ↔ is_reduced (R ⧸ I) := -by { conv_lhs { rw ← @ideal.mk_ker R _ I }, - exact ring_hom.ker_is_radical_iff_reduced_of_surjective (@ideal.quotient.mk_surjective R _ I) } - /-- An element `y` in a monoid is radical if for any element `x`, `y` divides `x` whenever it divides a power of `x`. -/ def is_radical [has_dvd R] [has_pow R ℕ] (y : R) : Prop := ∀ (n : ℕ) x, y ∣ x ^ n → y ∣ x @@ -228,65 +223,3 @@ begin end end module.End - -section ideal - -variables [comm_semiring R] [comm_ring S] [algebra R S] (I : ideal S) - -/-- Let `P` be a property on ideals. If `P` holds for square-zero ideals, and if - `P I → P (J ⧸ I) → P J`, then `P` holds for all nilpotent ideals. -/ -lemma ideal.is_nilpotent.induction_on - (hI : is_nilpotent I) - {P : ∀ ⦃S : Type*⦄ [comm_ring S], by exactI ∀ I : ideal S, Prop} - (h₁ : ∀ ⦃S : Type*⦄ [comm_ring S], by exactI ∀ I : ideal S, I ^ 2 = ⊥ → P I) - (h₂ : ∀ ⦃S : Type*⦄ [comm_ring S], by exactI - ∀ I J : ideal S, I ≤ J → P I → P (J.map (ideal.quotient.mk I)) → P J) : P I := -begin - obtain ⟨n, hI : I ^ n = ⊥⟩ := hI, - unfreezingI { revert S }, - apply nat.strong_induction_on n, - clear n, - introsI n H S _ I hI, - by_cases hI' : I = ⊥, - { subst hI', apply h₁, rw [← ideal.zero_eq_bot, zero_pow], exact zero_lt_two }, - cases n, - { rw [pow_zero, ideal.one_eq_top] at hI, - haveI := subsingleton_of_bot_eq_top hI.symm, - exact (hI' (subsingleton.elim _ _)).elim }, - cases n, - { rw [pow_one] at hI, - exact (hI' hI).elim }, - apply h₂ (I ^ 2) _ (ideal.pow_le_self two_ne_zero), - { apply H n.succ _ (I ^ 2), - { rw [← pow_mul, eq_bot_iff, ← hI, nat.succ_eq_add_one, nat.succ_eq_add_one], - exact ideal.pow_le_pow (by linarith) }, - { exact le_refl n.succ.succ } }, - { apply h₁, rw [← ideal.map_pow, ideal.map_quotient_self] }, -end - -lemma is_nilpotent.is_unit_quotient_mk_iff {R : Type*} [comm_ring R] {I : ideal R} - (hI : is_nilpotent I) {x : R} : is_unit (ideal.quotient.mk I x) ↔ is_unit x := -begin - refine ⟨_, λ h, h.map I^.quotient.mk⟩, - revert x, - apply ideal.is_nilpotent.induction_on I hI; clear hI I, - swap, - { introv e h₁ h₂ h₃, - apply h₁, - apply h₂, - exactI h₃.map ((double_quot.quot_quot_equiv_quot_sup I J).trans - (ideal.quot_equiv_of_eq (sup_eq_right.mpr e))).symm.to_ring_hom }, - { introv e H, - resetI, - obtain ⟨y, hy⟩ := ideal.quotient.mk_surjective (↑(H.unit⁻¹) : S ⧸ I), - have : ideal.quotient.mk I (x * y) = ideal.quotient.mk I 1, - { rw [map_one, _root_.map_mul, hy, is_unit.mul_coe_inv] }, - rw ideal.quotient.eq at this, - have : (x * y - 1) ^ 2 = 0, - { rw [← ideal.mem_bot, ← e], exact ideal.pow_mem_pow this _ }, - have : x * (y * (2 - x * y)) = 1, - { rw [eq_comm, ← sub_eq_zero, ← this], ring }, - exact is_unit_of_mul_eq_one _ _ this } -end - -end ideal diff --git a/src/ring_theory/noetherian.lean b/src/ring_theory/noetherian.lean index 2e59f4a21cb44..612de749225cf 100644 --- a/src/ring_theory/noetherian.lean +++ b/src/ring_theory/noetherian.lean @@ -470,10 +470,6 @@ begin refine (submodule.restrict_scalars_embedding R S M).dual.well_founded h end -instance ideal.quotient.is_noetherian_ring {R : Type*} [comm_ring R] [h : is_noetherian_ring R] - (I : ideal R) : is_noetherian_ring (R ⧸ I) := -is_noetherian_ring_iff.mpr $ is_noetherian_of_tower R $ submodule.quotient.is_noetherian _ - theorem is_noetherian_of_fg_of_noetherian {R M} [ring R] [add_comm_group M] [module R M] (N : submodule R M) [is_noetherian_ring R] (hN : N.fg) : is_noetherian R N := let ⟨s, hs⟩ := hN in diff --git a/src/ring_theory/polynomial/basic.lean b/src/ring_theory/polynomial/basic.lean index 91c3f62e0b89d..e31e66cc6bad7 100644 --- a/src/ring_theory/polynomial/basic.lean +++ b/src/ring_theory/polynomial/basic.lean @@ -588,124 +588,56 @@ end ring section comm_ring variables [comm_ring R] -lemma quotient_map_C_eq_zero {I : ideal R} : - ∀ a ∈ I, ((quotient.mk (map (C : R →+* R[X]) I : ideal R[X])).comp C) a = 0 := -begin - intros a ha, - rw [ring_hom.comp_apply, quotient.eq_zero_iff_mem], - exact mem_map_of_mem _ ha, -end - -lemma eval₂_C_mk_eq_zero {I : ideal R} : - ∀ f ∈ (map (C : R →+* R[X]) I : ideal R[X]), eval₂_ring_hom (C.comp (quotient.mk I)) X f = 0 := -begin - intros a ha, - rw ← sum_monomial_eq a, - dsimp, - rw eval₂_sum, - refine finset.sum_eq_zero (λ n hn, _), - dsimp, - rw eval₂_monomial (C.comp (quotient.mk I)) X, - refine mul_eq_zero_of_left (polynomial.ext (λ m, _)) (X ^ n), - erw coeff_C, - by_cases h : m = 0, - { simpa [h] using quotient.eq_zero_iff_mem.2 ((mem_map_C_iff.1 ha) n) }, - { simp [h] } -end - -/-- If `I` is an ideal of `R`, then the ring polynomials over the quotient ring `I.quotient` is -isomorphic to the quotient of `R[X]` by the ideal `map C I`, -where `map C I` contains exactly the polynomials whose coefficients all lie in `I` -/ -def polynomial_quotient_equiv_quotient_polynomial (I : ideal R) : - (R ⧸ I)[X] ≃+* R[X] ⧸ (map C I : ideal R[X]) := -{ to_fun := eval₂_ring_hom - (quotient.lift I ((quotient.mk (map C I : ideal R[X])).comp C) quotient_map_C_eq_zero) - ((quotient.mk (map C I : ideal R[X]) X)), - inv_fun := quotient.lift (map C I : ideal R[X]) - (eval₂_ring_hom (C.comp (quotient.mk I)) X) eval₂_C_mk_eq_zero, - map_mul' := λ f g, by simp only [coe_eval₂_ring_hom, eval₂_mul], - map_add' := λ f g, by simp only [eval₂_add, coe_eval₂_ring_hom], - left_inv := begin - intro f, - apply polynomial.induction_on' f, - { intros p q hp hq, - simp only [coe_eval₂_ring_hom] at hp, - simp only [coe_eval₂_ring_hom] at hq, - simp only [coe_eval₂_ring_hom, hp, hq, ring_hom.map_add] }, - { rintros n ⟨x⟩, - simp only [← smul_X_eq_monomial, C_mul', quotient.lift_mk, submodule.quotient.quot_mk_eq_mk, - quotient.mk_eq_mk, eval₂_X_pow, eval₂_smul, coe_eval₂_ring_hom, ring_hom.map_pow, - eval₂_C, ring_hom.coe_comp, ring_hom.map_mul, eval₂_X] } - end, - right_inv := begin - rintro ⟨f⟩, - apply polynomial.induction_on' f, - { simp_intros p q hp hq, - rw [hp, hq] }, - { intros n a, - simp only [← smul_X_eq_monomial, ← C_mul' a (X ^ n), quotient.lift_mk, - submodule.quotient.quot_mk_eq_mk, quotient.mk_eq_mk, eval₂_X_pow, - eval₂_smul, coe_eval₂_ring_hom, ring_hom.map_pow, eval₂_C, ring_hom.coe_comp, - ring_hom.map_mul, eval₂_X] }, - end, } - -@[simp] -lemma polynomial_quotient_equiv_quotient_polynomial_symm_mk (I : ideal R) (f : R[X]) : - I.polynomial_quotient_equiv_quotient_polynomial.symm (quotient.mk _ f) = f.map (quotient.mk I) := -by rw [polynomial_quotient_equiv_quotient_polynomial, ring_equiv.symm_mk, ring_equiv.coe_mk, - ideal.quotient.lift_mk, coe_eval₂_ring_hom, eval₂_eq_eval_map, ←polynomial.map_map, - ←eval₂_eq_eval_map, polynomial.eval₂_C_X] - -@[simp] -lemma polynomial_quotient_equiv_quotient_polynomial_map_mk (I : ideal R) (f : R[X]) : - I.polynomial_quotient_equiv_quotient_polynomial (f.map I^.quotient.mk) = quotient.mk _ f := +/-- If `P` is a prime ideal of `R`, then `P.R[x]` is a prime ideal of `R[x]`. -/ +lemma is_prime_map_C_iff_is_prime (P : ideal R) : + is_prime (map (C : R →+* R[X]) P : ideal R[X]) ↔ is_prime P := begin - apply (polynomial_quotient_equiv_quotient_polynomial I).symm.injective, - rw [ring_equiv.symm_apply_apply, polynomial_quotient_equiv_quotient_polynomial_symm_mk], + -- Porting note: the following proof avoids quotient rings + -- It can be golfed substantially by using something like + -- `(quotient.is_domain_iff_prime (map C P : ideal R[X]))` + split, + { intro H, + have := @comap_is_prime R R[X] (R →+* R[X]) _ _ _ C (map C P) H, + convert this using 1, + ext x, + simp only [mem_comap, mem_map_C_iff], + split, + { rintro h (-|n), + { simpa only [coeff_C_zero] using h }, + { simp only [coeff_C_ne_zero (nat.succ_ne_zero _), submodule.zero_mem] } }, + { intro h, simpa only [coeff_C_zero] using h 0 } }, + { intro h, + constructor, + { rw [ne.def, eq_top_iff_one, mem_map_C_iff, not_forall], + use 0, + rw [coeff_one_zero, ← eq_top_iff_one], exact h.1 }, + { intros f g, simp only [mem_map_C_iff], contrapose!, + rintro ⟨hf, hg⟩, + classical, + let m := nat.find hf, + let n := nat.find hg, + refine ⟨m + n, _⟩, + rw [coeff_mul, ← finset.insert_erase ((@finset.nat.mem_antidiagonal _ (m,n)).mpr rfl), + finset.sum_insert (finset.not_mem_erase _ _), (P.add_mem_iff_left _).not], + { apply mt h.2, rw [not_or_distrib], exact ⟨nat.find_spec hf, nat.find_spec hg⟩ }, + apply P.sum_mem, + rintro ⟨i, j⟩ hij, + rw [finset.mem_erase, finset.nat.mem_antidiagonal] at hij, + simp only [ne.def, prod.mk.inj_iff, not_and_distrib] at hij, + obtain (hi|hj) : i < m ∨ j < n, + { rw [or_iff_not_imp_left, not_lt, le_iff_lt_or_eq], + rintro (hmi|rfl), + { rw [← not_le], intro hnj, exact (add_lt_add_of_lt_of_le hmi hnj).ne hij.2.symm, }, + { simpa only [eq_self_iff_true, not_true, false_or, add_right_inj, not_and_self] + using hij, } }, + { rw [mul_comm], apply P.mul_mem_left, exact not_not.1 (nat.find_min hf hi) }, + { apply P.mul_mem_left, exact not_not.1 (nat.find_min hg hj) } } } end -/-- If `P` is a prime ideal of `R`, then `R[x]/(P)` is an integral domain. -/ -lemma is_domain_map_C_quotient {P : ideal R} (H : is_prime P) : - is_domain (R[X] ⧸ (map (C : R →+* R[X]) P : ideal R[X])) := -ring_equiv.is_domain (polynomial (R ⧸ P)) - (polynomial_quotient_equiv_quotient_polynomial P).symm - /-- If `P` is a prime ideal of `R`, then `P.R[x]` is a prime ideal of `R[x]`. -/ lemma is_prime_map_C_of_is_prime {P : ideal R} (H : is_prime P) : is_prime (map (C : R →+* R[X]) P : ideal R[X]) := -(quotient.is_domain_iff_prime (map C P : ideal R[X])).mp - (is_domain_map_C_quotient H) - -/-- Given any ring `R` and an ideal `I` of `R[X]`, we get a map `R → R[x] → R[x]/I`. - If we let `R` be the image of `R` in `R[x]/I` then we also have a map `R[x] → R'[x]`. - In particular we can map `I` across this map, to get `I'` and a new map `R' → R'[x] → R'[x]/I`. - This theorem shows `I'` will not contain any non-zero constant polynomials - -/ -lemma eq_zero_of_polynomial_mem_map_range (I : ideal R[X]) - (x : ((quotient.mk I).comp C).range) - (hx : C x ∈ (I.map (polynomial.map_ring_hom ((quotient.mk I).comp C).range_restrict))) : - x = 0 := -begin - let i := ((quotient.mk I).comp C).range_restrict, - have hi' : (polynomial.map_ring_hom i).ker ≤ I, - { refine λ f hf, polynomial_mem_ideal_of_coeff_mem_ideal I f (λ n, _), - rw [mem_comap, ← quotient.eq_zero_iff_mem, ← ring_hom.comp_apply], - rw [ring_hom.mem_ker, coe_map_ring_hom] at hf, - replace hf := congr_arg (λ (f : polynomial _), f.coeff n) hf, - simp only [coeff_map, coeff_zero] at hf, - rwa [subtype.ext_iff, ring_hom.coe_range_restrict] at hf }, - obtain ⟨x, hx'⟩ := x, - obtain ⟨y, rfl⟩ := (ring_hom.mem_range).1 hx', - refine subtype.eq _, - simp only [ring_hom.comp_apply, quotient.eq_zero_iff_mem, zero_mem_class.coe_zero, - subtype.val_eq_coe], - suffices : C (i y) ∈ (I.map (polynomial.map_ring_hom i)), - { obtain ⟨f, hf⟩ := mem_image_of_mem_map_of_surjective (polynomial.map_ring_hom i) - (polynomial.map_surjective _ (((quotient.mk I).comp C).range_restrict_surjective)) this, - refine sub_add_cancel (C y) f ▸ I.add_mem (hi' _ : (C y - f) ∈ I) hf.1, - rw [ring_hom.mem_ker, ring_hom.map_sub, hf.2, sub_eq_zero, coe_map_ring_hom, map_C] }, - exact hx, -end +(is_prime_map_C_iff_is_prime P).mpr H theorem is_fg_degree_le [is_noetherian_ring R] (I : ideal R[X]) (n : ℕ) : submodule.fg (I.degree_le n) := @@ -1057,14 +989,6 @@ begin simp only [monomial_eq, ϕ.map_pow, ϕ.map_prod, ϕ.comp_apply, ϕ.map_mul, finsupp.prod_pow], end -lemma quotient_map_C_eq_zero {I : ideal R} {i : R} (hi : i ∈ I) : - (ideal.quotient.mk (ideal.map (C : R →+* mv_polynomial σ R) I : - ideal (mv_polynomial σ R))).comp C i = 0 := -begin - simp only [function.comp_app, ring_hom.coe_comp, ideal.quotient.eq_zero_iff_mem], - exact ideal.mem_map_of_mem _ hi -end - /-- If every coefficient of a polynomial is in an ideal `I`, then so is the polynomial itself, multivariate version. -/ lemma mem_ideal_of_coeff_mem_ideal (I : ideal (mv_polynomial σ R)) (p : mv_polynomial σ R) @@ -1121,63 +1045,6 @@ begin simp_rw [coeff_map, coeff_zero, ring_hom.mem_ker], end -lemma eval₂_C_mk_eq_zero {I : ideal R} {a : mv_polynomial σ R} - (ha : a ∈ (ideal.map (C : R →+* mv_polynomial σ R) I : ideal (mv_polynomial σ R))) : - eval₂_hom (C.comp (ideal.quotient.mk I)) X a = 0 := -begin - rw as_sum a, - rw [coe_eval₂_hom, eval₂_sum], - refine finset.sum_eq_zero (λ n hn, _), - simp only [eval₂_monomial, function.comp_app, ring_hom.coe_comp], - refine mul_eq_zero_of_left _ _, - suffices : coeff n a ∈ I, - { rw [← @ideal.mk_ker R _ I, ring_hom.mem_ker] at this, - simp only [this, C_0] }, - exact mem_map_C_iff.1 ha n -end - -/-- If `I` is an ideal of `R`, then the ring `mv_polynomial σ I.quotient` is isomorphic as an -`R`-algebra to the quotient of `mv_polynomial σ R` by the ideal generated by `I`. -/ -def quotient_equiv_quotient_mv_polynomial (I : ideal R) : - mv_polynomial σ (R ⧸ I) ≃ₐ[R] - mv_polynomial σ R ⧸ (ideal.map C I : ideal (mv_polynomial σ R)) := -{ to_fun := eval₂_hom (ideal.quotient.lift I ((ideal.quotient.mk (ideal.map C I : ideal - (mv_polynomial σ R))).comp C) (λ i hi, quotient_map_C_eq_zero hi)) - (λ i, ideal.quotient.mk (ideal.map C I : ideal (mv_polynomial σ R)) (X i)), - inv_fun := ideal.quotient.lift (ideal.map C I : ideal (mv_polynomial σ R)) - (eval₂_hom (C.comp (ideal.quotient.mk I)) X) (λ a ha, eval₂_C_mk_eq_zero ha), - map_mul' := ring_hom.map_mul _, - map_add' := ring_hom.map_add _, - left_inv := begin - intro f, - apply induction_on f, - { rintro ⟨r⟩, - rw [coe_eval₂_hom, eval₂_C], - simp only [eval₂_hom_eq_bind₂, submodule.quotient.quot_mk_eq_mk, ideal.quotient.lift_mk, - ideal.quotient.mk_eq_mk, bind₂_C_right, ring_hom.coe_comp] }, - { simp_intros p q hp hq only [ring_hom.map_add, mv_polynomial.coe_eval₂_hom, coe_eval₂_hom, - mv_polynomial.eval₂_add, mv_polynomial.eval₂_hom_eq_bind₂, eval₂_hom_eq_bind₂], - rw [hp, hq] }, - { simp_intros p i hp only [eval₂_hom_eq_bind₂, coe_eval₂_hom], - simp only [hp, eval₂_hom_eq_bind₂, coe_eval₂_hom, ideal.quotient.lift_mk, bind₂_X_right, - eval₂_mul, ring_hom.map_mul, eval₂_X] } - end, - right_inv := begin - rintro ⟨f⟩, - apply induction_on f, - { intros r, - simp only [submodule.quotient.quot_mk_eq_mk, ideal.quotient.lift_mk, ideal.quotient.mk_eq_mk, - ring_hom.coe_comp, eval₂_hom_C] }, - { simp_intros p q hp hq only [eval₂_hom_eq_bind₂, submodule.quotient.quot_mk_eq_mk, eval₂_add, - ring_hom.map_add, coe_eval₂_hom, ideal.quotient.lift_mk, ideal.quotient.mk_eq_mk], - rw [hp, hq] }, - { simp_intros p i hp only [eval₂_hom_eq_bind₂, submodule.quotient.quot_mk_eq_mk, coe_eval₂_hom, - ideal.quotient.lift_mk, ideal.quotient.mk_eq_mk, bind₂_X_right, eval₂_mul, ring_hom.map_mul, - eval₂_X], - simp only [hp] } - end, - commutes' := λ r, eval₂_hom_C _ _ (ideal.quotient.mk I r) } - end mv_polynomial section unique_factorization_domain diff --git a/src/ring_theory/polynomial/quotient.lean b/src/ring_theory/polynomial/quotient.lean new file mode 100644 index 0000000000000..155a1a2c368c8 --- /dev/null +++ b/src/ring_theory/polynomial/quotient.lean @@ -0,0 +1,229 @@ +/- +Copyright (c) 2019 Kenny Lau. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Kenny Lau, David Kurniadi Angdinata, Devon Tuma, Riccardo Brasca +-/ + +import data.polynomial.div +import ring_theory.polynomial.basic +import ring_theory.ideal.quotient_operations + +/-! +# Quotients of polynomial rings +-/ + +open_locale polynomial + +namespace polynomial + +variables {R : Type*} [comm_ring R] + +/-- For a commutative ring $R$, evaluating a polynomial at an element $x \in R$ induces an +isomorphism of $R$-algebras $R[X] / \langle X - x \rangle \cong R$. -/ +noncomputable def quotient_span_X_sub_C_alg_equiv (x : R) : + (R[X] ⧸ ideal.span ({X - C x} : set R[X])) ≃ₐ[R] R := +(alg_equiv.restrict_scalars R $ ideal.quotient_equiv_alg_of_eq R + (by exact ker_eval_ring_hom x : ring_hom.ker (aeval x).to_ring_hom = _)).symm.trans $ + ideal.quotient_ker_alg_equiv_of_right_inverse $ λ _, eval_C + +@[simp] lemma quotient_span_X_sub_C_alg_equiv_mk (x : R) (p : R[X]) : + quotient_span_X_sub_C_alg_equiv x (ideal.quotient.mk _ p) = p.eval x := +rfl + +@[simp] lemma quotient_span_X_sub_C_alg_equiv_symm_apply (x : R) (y : R) : + (quotient_span_X_sub_C_alg_equiv x).symm y = algebra_map R _ y := +rfl + +end polynomial + +namespace ideal +noncomputable theory +open polynomial + +variables {R : Type*} [comm_ring R] + +lemma quotient_map_C_eq_zero {I : ideal R} : + ∀ a ∈ I, ((quotient.mk (map (C : R →+* R[X]) I : ideal R[X])).comp C) a = 0 := +begin + intros a ha, + rw [ring_hom.comp_apply, quotient.eq_zero_iff_mem], + exact mem_map_of_mem _ ha, +end + +lemma eval₂_C_mk_eq_zero {I : ideal R} : + ∀ f ∈ (map (C : R →+* R[X]) I : ideal R[X]), eval₂_ring_hom (C.comp (quotient.mk I)) X f = 0 := +begin + intros a ha, + rw ← sum_monomial_eq a, + dsimp, + rw eval₂_sum, + refine finset.sum_eq_zero (λ n hn, _), + dsimp, + rw eval₂_monomial (C.comp (quotient.mk I)) X, + refine mul_eq_zero_of_left (polynomial.ext (λ m, _)) (X ^ n), + erw coeff_C, + by_cases h : m = 0, + { simpa [h] using quotient.eq_zero_iff_mem.2 ((mem_map_C_iff.1 ha) n) }, + { simp [h] } +end + +/-- If `I` is an ideal of `R`, then the ring polynomials over the quotient ring `I.quotient` is +isomorphic to the quotient of `R[X]` by the ideal `map C I`, +where `map C I` contains exactly the polynomials whose coefficients all lie in `I` -/ +def polynomial_quotient_equiv_quotient_polynomial (I : ideal R) : + (R ⧸ I)[X] ≃+* R[X] ⧸ (map C I : ideal R[X]) := +{ to_fun := eval₂_ring_hom + (quotient.lift I ((quotient.mk (map C I : ideal R[X])).comp C) quotient_map_C_eq_zero) + ((quotient.mk (map C I : ideal R[X]) X)), + inv_fun := quotient.lift (map C I : ideal R[X]) + (eval₂_ring_hom (C.comp (quotient.mk I)) X) eval₂_C_mk_eq_zero, + map_mul' := λ f g, by simp only [coe_eval₂_ring_hom, eval₂_mul], + map_add' := λ f g, by simp only [eval₂_add, coe_eval₂_ring_hom], + left_inv := begin + intro f, + apply polynomial.induction_on' f, + { intros p q hp hq, + simp only [coe_eval₂_ring_hom] at hp, + simp only [coe_eval₂_ring_hom] at hq, + simp only [coe_eval₂_ring_hom, hp, hq, ring_hom.map_add] }, + { rintros n ⟨x⟩, + simp only [← smul_X_eq_monomial, C_mul', quotient.lift_mk, submodule.quotient.quot_mk_eq_mk, + quotient.mk_eq_mk, eval₂_X_pow, eval₂_smul, coe_eval₂_ring_hom, ring_hom.map_pow, + eval₂_C, ring_hom.coe_comp, ring_hom.map_mul, eval₂_X] } + end, + right_inv := begin + rintro ⟨f⟩, + apply polynomial.induction_on' f, + { simp_intros p q hp hq, + rw [hp, hq] }, + { intros n a, + simp only [← smul_X_eq_monomial, ← C_mul' a (X ^ n), quotient.lift_mk, + submodule.quotient.quot_mk_eq_mk, quotient.mk_eq_mk, eval₂_X_pow, + eval₂_smul, coe_eval₂_ring_hom, ring_hom.map_pow, eval₂_C, ring_hom.coe_comp, + ring_hom.map_mul, eval₂_X] }, + end, } + +@[simp] +lemma polynomial_quotient_equiv_quotient_polynomial_symm_mk (I : ideal R) (f : R[X]) : + I.polynomial_quotient_equiv_quotient_polynomial.symm (quotient.mk _ f) = f.map (quotient.mk I) := +by rw [polynomial_quotient_equiv_quotient_polynomial, ring_equiv.symm_mk, ring_equiv.coe_mk, + ideal.quotient.lift_mk, coe_eval₂_ring_hom, eval₂_eq_eval_map, ←polynomial.map_map, + ←eval₂_eq_eval_map, polynomial.eval₂_C_X] + +@[simp] +lemma polynomial_quotient_equiv_quotient_polynomial_map_mk (I : ideal R) (f : R[X]) : + I.polynomial_quotient_equiv_quotient_polynomial (f.map I^.quotient.mk) = quotient.mk _ f := +begin + apply (polynomial_quotient_equiv_quotient_polynomial I).symm.injective, + rw [ring_equiv.symm_apply_apply, polynomial_quotient_equiv_quotient_polynomial_symm_mk], +end + +/-- If `P` is a prime ideal of `R`, then `R[x]/(P)` is an integral domain. -/ +lemma is_domain_map_C_quotient {P : ideal R} (H : is_prime P) : + is_domain (R[X] ⧸ (map (C : R →+* R[X]) P : ideal R[X])) := +ring_equiv.is_domain (polynomial (R ⧸ P)) + (polynomial_quotient_equiv_quotient_polynomial P).symm + +/-- Given any ring `R` and an ideal `I` of `R[X]`, we get a map `R → R[x] → R[x]/I`. + If we let `R` be the image of `R` in `R[x]/I` then we also have a map `R[x] → R'[x]`. + In particular we can map `I` across this map, to get `I'` and a new map `R' → R'[x] → R'[x]/I`. + This theorem shows `I'` will not contain any non-zero constant polynomials + -/ +lemma eq_zero_of_polynomial_mem_map_range (I : ideal R[X]) + (x : ((quotient.mk I).comp C).range) + (hx : C x ∈ (I.map (polynomial.map_ring_hom ((quotient.mk I).comp C).range_restrict))) : + x = 0 := +begin + let i := ((quotient.mk I).comp C).range_restrict, + have hi' : (polynomial.map_ring_hom i).ker ≤ I, + { refine λ f hf, polynomial_mem_ideal_of_coeff_mem_ideal I f (λ n, _), + rw [mem_comap, ← quotient.eq_zero_iff_mem, ← ring_hom.comp_apply], + rw [ring_hom.mem_ker, coe_map_ring_hom] at hf, + replace hf := congr_arg (λ (f : polynomial _), f.coeff n) hf, + simp only [coeff_map, coeff_zero] at hf, + rwa [subtype.ext_iff, ring_hom.coe_range_restrict] at hf }, + obtain ⟨x, hx'⟩ := x, + obtain ⟨y, rfl⟩ := (ring_hom.mem_range).1 hx', + refine subtype.eq _, + simp only [ring_hom.comp_apply, quotient.eq_zero_iff_mem, zero_mem_class.coe_zero, + subtype.val_eq_coe], + suffices : C (i y) ∈ (I.map (polynomial.map_ring_hom i)), + { obtain ⟨f, hf⟩ := mem_image_of_mem_map_of_surjective (polynomial.map_ring_hom i) + (polynomial.map_surjective _ (((quotient.mk I).comp C).range_restrict_surjective)) this, + refine sub_add_cancel (C y) f ▸ I.add_mem (hi' _ : (C y - f) ∈ I) hf.1, + rw [ring_hom.mem_ker, ring_hom.map_sub, hf.2, sub_eq_zero, coe_map_ring_hom, map_C] }, + exact hx, +end + +end ideal + +namespace mv_polynomial + +variables {R : Type*} {σ : Type*} [comm_ring R] {r : R} + +lemma quotient_map_C_eq_zero {I : ideal R} {i : R} (hi : i ∈ I) : + (ideal.quotient.mk (ideal.map (C : R →+* mv_polynomial σ R) I : + ideal (mv_polynomial σ R))).comp C i = 0 := +begin + simp only [function.comp_app, ring_hom.coe_comp, ideal.quotient.eq_zero_iff_mem], + exact ideal.mem_map_of_mem _ hi +end + +lemma eval₂_C_mk_eq_zero {I : ideal R} {a : mv_polynomial σ R} + (ha : a ∈ (ideal.map (C : R →+* mv_polynomial σ R) I : ideal (mv_polynomial σ R))) : + eval₂_hom (C.comp (ideal.quotient.mk I)) X a = 0 := +begin + rw as_sum a, + rw [coe_eval₂_hom, eval₂_sum], + refine finset.sum_eq_zero (λ n hn, _), + simp only [eval₂_monomial, function.comp_app, ring_hom.coe_comp], + refine mul_eq_zero_of_left _ _, + suffices : coeff n a ∈ I, + { rw [← @ideal.mk_ker R _ I, ring_hom.mem_ker] at this, + simp only [this, C_0] }, + exact mem_map_C_iff.1 ha n +end + +/-- If `I` is an ideal of `R`, then the ring `mv_polynomial σ I.quotient` is isomorphic as an +`R`-algebra to the quotient of `mv_polynomial σ R` by the ideal generated by `I`. -/ +def quotient_equiv_quotient_mv_polynomial (I : ideal R) : + mv_polynomial σ (R ⧸ I) ≃ₐ[R] + mv_polynomial σ R ⧸ (ideal.map C I : ideal (mv_polynomial σ R)) := +{ to_fun := eval₂_hom (ideal.quotient.lift I ((ideal.quotient.mk (ideal.map C I : ideal + (mv_polynomial σ R))).comp C) (λ i hi, quotient_map_C_eq_zero hi)) + (λ i, ideal.quotient.mk (ideal.map C I : ideal (mv_polynomial σ R)) (X i)), + inv_fun := ideal.quotient.lift (ideal.map C I : ideal (mv_polynomial σ R)) + (eval₂_hom (C.comp (ideal.quotient.mk I)) X) (λ a ha, eval₂_C_mk_eq_zero ha), + map_mul' := ring_hom.map_mul _, + map_add' := ring_hom.map_add _, + left_inv := begin + intro f, + apply induction_on f, + { rintro ⟨r⟩, + rw [coe_eval₂_hom, eval₂_C], + simp only [eval₂_hom_eq_bind₂, submodule.quotient.quot_mk_eq_mk, ideal.quotient.lift_mk, + ideal.quotient.mk_eq_mk, bind₂_C_right, ring_hom.coe_comp] }, + { simp_intros p q hp hq only [ring_hom.map_add, mv_polynomial.coe_eval₂_hom, coe_eval₂_hom, + mv_polynomial.eval₂_add, mv_polynomial.eval₂_hom_eq_bind₂, eval₂_hom_eq_bind₂], + rw [hp, hq] }, + { simp_intros p i hp only [eval₂_hom_eq_bind₂, coe_eval₂_hom], + simp only [hp, eval₂_hom_eq_bind₂, coe_eval₂_hom, ideal.quotient.lift_mk, bind₂_X_right, + eval₂_mul, ring_hom.map_mul, eval₂_X] } + end, + right_inv := begin + rintro ⟨f⟩, + apply induction_on f, + { intros r, + simp only [submodule.quotient.quot_mk_eq_mk, ideal.quotient.lift_mk, ideal.quotient.mk_eq_mk, + ring_hom.coe_comp, eval₂_hom_C] }, + { simp_intros p q hp hq only [eval₂_hom_eq_bind₂, submodule.quotient.quot_mk_eq_mk, eval₂_add, + ring_hom.map_add, coe_eval₂_hom, ideal.quotient.lift_mk, ideal.quotient.mk_eq_mk], + rw [hp, hq] }, + { simp_intros p i hp only [eval₂_hom_eq_bind₂, submodule.quotient.quot_mk_eq_mk, coe_eval₂_hom, + ideal.quotient.lift_mk, ideal.quotient.mk_eq_mk, bind₂_X_right, eval₂_mul, ring_hom.map_mul, + eval₂_X], + simp only [hp] } + end, + commutes' := λ r, eval₂_hom_C _ _ (ideal.quotient.mk I r) } + +end mv_polynomial diff --git a/src/ring_theory/quotient_nilpotent.lean b/src/ring_theory/quotient_nilpotent.lean new file mode 100644 index 0000000000000..95eb145876376 --- /dev/null +++ b/src/ring_theory/quotient_nilpotent.lean @@ -0,0 +1,74 @@ +/- +Copyright (c) 2021 Oliver Nash. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Oliver Nash +-/ +import ring_theory.nilpotent +import ring_theory.ideal.quotient_operations + +/-! +# Nilpotent elements in quotient rings +-/ + +lemma ideal.is_radical_iff_quotient_reduced {R : Type*} [comm_ring R] (I : ideal R) : + I.is_radical ↔ is_reduced (R ⧸ I) := +by { conv_lhs { rw ← @ideal.mk_ker R _ I }, + exact ring_hom.ker_is_radical_iff_reduced_of_surjective (@ideal.quotient.mk_surjective R _ I) } + +variables {R S : Type*} [comm_semiring R] [comm_ring S] [algebra R S] (I : ideal S) + +/-- Let `P` be a property on ideals. If `P` holds for square-zero ideals, and if + `P I → P (J ⧸ I) → P J`, then `P` holds for all nilpotent ideals. -/ +lemma ideal.is_nilpotent.induction_on + (hI : is_nilpotent I) + {P : ∀ ⦃S : Type*⦄ [comm_ring S], by exactI ∀ I : ideal S, Prop} + (h₁ : ∀ ⦃S : Type*⦄ [comm_ring S], by exactI ∀ I : ideal S, I ^ 2 = ⊥ → P I) + (h₂ : ∀ ⦃S : Type*⦄ [comm_ring S], by exactI + ∀ I J : ideal S, I ≤ J → P I → P (J.map (ideal.quotient.mk I)) → P J) : P I := +begin + obtain ⟨n, hI : I ^ n = ⊥⟩ := hI, + unfreezingI { revert S }, + apply nat.strong_induction_on n, + clear n, + introsI n H S _ I hI, + by_cases hI' : I = ⊥, + { subst hI', apply h₁, rw [← ideal.zero_eq_bot, zero_pow], exact zero_lt_two }, + cases n, + { rw [pow_zero, ideal.one_eq_top] at hI, + haveI := subsingleton_of_bot_eq_top hI.symm, + exact (hI' (subsingleton.elim _ _)).elim }, + cases n, + { rw [pow_one] at hI, + exact (hI' hI).elim }, + apply h₂ (I ^ 2) _ (ideal.pow_le_self two_ne_zero), + { apply H n.succ _ (I ^ 2), + { rw [← pow_mul, eq_bot_iff, ← hI, nat.succ_eq_add_one, nat.succ_eq_add_one], + exact ideal.pow_le_pow (by linarith) }, + { exact le_refl n.succ.succ } }, + { apply h₁, rw [← ideal.map_pow, ideal.map_quotient_self] }, +end + +lemma is_nilpotent.is_unit_quotient_mk_iff {R : Type*} [comm_ring R] {I : ideal R} + (hI : is_nilpotent I) {x : R} : is_unit (ideal.quotient.mk I x) ↔ is_unit x := +begin + refine ⟨_, λ h, h.map I^.quotient.mk⟩, + revert x, + apply ideal.is_nilpotent.induction_on I hI; clear hI I, + swap, + { introv e h₁ h₂ h₃, + apply h₁, + apply h₂, + exactI h₃.map ((double_quot.quot_quot_equiv_quot_sup I J).trans + (ideal.quot_equiv_of_eq (sup_eq_right.mpr e))).symm.to_ring_hom }, + { introv e H, + resetI, + obtain ⟨y, hy⟩ := ideal.quotient.mk_surjective (↑(H.unit⁻¹) : S ⧸ I), + have : ideal.quotient.mk I (x * y) = ideal.quotient.mk I 1, + { rw [map_one, _root_.map_mul, hy, is_unit.mul_coe_inv] }, + rw ideal.quotient.eq at this, + have : (x * y - 1) ^ 2 = 0, + { rw [← ideal.mem_bot, ← e], exact ideal.pow_mem_pow this _ }, + have : x * (y * (2 - x * y)) = 1, + { rw [eq_comm, ← sub_eq_zero, ← this], ring }, + exact is_unit_of_mul_eq_one _ _ this } +end diff --git a/src/ring_theory/quotient_noetherian.lean b/src/ring_theory/quotient_noetherian.lean new file mode 100644 index 0000000000000..b430736ab1ee8 --- /dev/null +++ b/src/ring_theory/quotient_noetherian.lean @@ -0,0 +1,15 @@ +/- +Copyright (c) 2021 Anne Baanen. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Anne Baanen +-/ +import ring_theory.noetherian +import ring_theory.quotient_nilpotent + +/-! +# Noetherian quotient rings and quotient modules +-/ + +instance ideal.quotient.is_noetherian_ring {R : Type*} [comm_ring R] [h : is_noetherian_ring R] + (I : ideal R) : is_noetherian_ring (R ⧸ I) := +is_noetherian_ring_iff.mpr $ is_noetherian_of_tower R $ submodule.quotient.is_noetherian _ diff --git a/src/ring_theory/valuation/basic.lean b/src/ring_theory/valuation/basic.lean index 82adda471a5d2..b3859e9586977 100644 --- a/src/ring_theory/valuation/basic.lean +++ b/src/ring_theory/valuation/basic.lean @@ -5,7 +5,7 @@ Authors: Kevin Buzzard, Johan Commelin, Patrick Massot -/ import algebra.order.with_zero -import ring_theory.ideal.quotient_operations +import ring_theory.ideal.operations /-! @@ -31,10 +31,6 @@ that the class of all valuations (as `Γ₀` varies) on a ring `R` is not a set. The "relation" is however reflexive, symmetric and transitive in the obvious sense. Note that we use 1.27(iii) of [wedhorn_adic] as the definition of equivalence. -The support of a valuation `v : valuation R Γ₀` is `supp v`. If `J` is an ideal of `R` -with `h : J ⊆ supp v` then the induced valuation -on R / J = `ideal.quotient J` is `on_quot v h`. - ## Main definitions * `valuation R Γ₀`, the type of valuations on `R` with values in `Γ₀` @@ -504,28 +500,6 @@ begin ... ≤ v (a + s) : aux (a + s) (-s) (by rwa ←ideal.neg_mem_iff at h) end -/-- If `hJ : J ⊆ supp v` then `on_quot_val hJ` is the induced function on R/J as a function. -Note: it's just the function; the valuation is `on_quot hJ`. -/ -def on_quot_val {J : ideal R} (hJ : J ≤ supp v) : - R ⧸ J → Γ₀ := -λ q, quotient.lift_on' q v $ λ a b h, -calc v a = v (b + -(-a + b)) : by simp - ... = v b : - v.map_add_supp b $ (ideal.neg_mem_iff _).2 $ hJ $ quotient_add_group.left_rel_apply.mp h - -/-- The extension of valuation v on R to valuation on R/J if J ⊆ supp v -/ -def on_quot {J : ideal R} (hJ : J ≤ supp v) : - valuation (R ⧸ J) Γ₀ := -{ to_fun := v.on_quot_val hJ, - map_zero' := v.map_zero, - map_one' := v.map_one, - map_mul' := λ xbar ybar, quotient.ind₂' v.map_mul xbar ybar, - map_add_le_max' := λ xbar ybar, quotient.ind₂' v.map_add xbar ybar } - -@[simp] lemma on_quot_comap_eq {J : ideal R} (hJ : J ≤ supp v) : - (v.on_quot hJ).comap (ideal.quotient.mk J) = v := -ext $ λ r, rfl - lemma comap_supp {S : Type*} [comm_ring S] (f : S →+* R) : supp (v.comap f) = ideal.comap f v.supp := ideal.ext $ λ x, @@ -534,29 +508,6 @@ begin refl, end -lemma self_le_supp_comap (J : ideal R) (v : valuation (R ⧸ J) Γ₀) : - J ≤ (v.comap (ideal.quotient.mk J)).supp := -by { rw [comap_supp, ← ideal.map_le_iff_le_comap], simp } - -@[simp] lemma comap_on_quot_eq (J : ideal R) (v : valuation (R ⧸ J) Γ₀) : - (v.comap (ideal.quotient.mk J)).on_quot (v.self_le_supp_comap J) = v := -ext $ by { rintro ⟨x⟩, refl } - -/-- The quotient valuation on R/J has support supp(v)/J if J ⊆ supp v. -/ -lemma supp_quot {J : ideal R} (hJ : J ≤ supp v) : - supp (v.on_quot hJ) = (supp v).map (ideal.quotient.mk J) := -begin - apply le_antisymm, - { rintro ⟨x⟩ hx, - apply ideal.subset_span, - exact ⟨x, hx, rfl⟩ }, - { rw ideal.map_le_iff_le_comap, - intros x hx, exact hx } -end - -lemma supp_quot_supp : supp (v.on_quot le_rfl) = 0 := -by { rw supp_quot, exact ideal.map_quotient_self _ } - end supp -- end of section end valuation @@ -766,39 +717,6 @@ def supp : ideal R := v.supp lemma map_add_supp (a : R) {s : R} (h : s ∈ supp v) : v (a + s) = v a := v.map_add_supp a h -/-- If `hJ : J ⊆ supp v` then `on_quot_val hJ` is the induced function on R/J as a function. -Note: it's just the function; the valuation is `on_quot hJ`. -/ -def on_quot_val {J : ideal R} (hJ : J ≤ supp v) : (R ⧸ J) → Γ₀ := v.on_quot_val hJ - -/-- The extension of valuation v on R to valuation on R/J if J ⊆ supp v -/ -def on_quot {J : ideal R} (hJ : J ≤ supp v) : - add_valuation (R ⧸ J) Γ₀ := -v.on_quot hJ - -@[simp] lemma on_quot_comap_eq {J : ideal R} (hJ : J ≤ supp v) : - (v.on_quot hJ).comap (ideal.quotient.mk J) = v := -v.on_quot_comap_eq hJ - -lemma comap_supp {S : Type*} [comm_ring S] (f : S →+* R) : - supp (v.comap f) = ideal.comap f v.supp := -v.comap_supp f - -lemma self_le_supp_comap (J : ideal R) (v : add_valuation (R ⧸ J) Γ₀) : - J ≤ (v.comap (ideal.quotient.mk J)).supp := -v.self_le_supp_comap J - -@[simp] lemma comap_on_quot_eq (J : ideal R) (v : add_valuation (R ⧸ J) Γ₀) : - (v.comap (ideal.quotient.mk J)).on_quot (v.self_le_supp_comap J) = v := -v.comap_on_quot_eq J - -/-- The quotient valuation on R/J has support supp(v)/J if J ⊆ supp v. -/ -lemma supp_quot {J : ideal R} (hJ : J ≤ supp v) : - supp (v.on_quot hJ) = (supp v).map (ideal.quotient.mk J) := -v.supp_quot hJ - -lemma supp_quot_supp : supp (v.on_quot le_rfl) = 0 := -v.supp_quot_supp - end supp -- end of section attribute [irreducible] add_valuation diff --git a/src/ring_theory/valuation/quotient.lean b/src/ring_theory/valuation/quotient.lean new file mode 100644 index 0000000000000..dde2f37caf667 --- /dev/null +++ b/src/ring_theory/valuation/quotient.lean @@ -0,0 +1,112 @@ +/- +Copyright (c) 2020 Johan Commelin. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Kevin Buzzard, Johan Commelin, Patrick Massot +-/ + +import ring_theory.valuation.basic +import ring_theory.ideal.quotient_operations + +/-! +# The valuation on a quotient ring + +The support of a valuation `v : valuation R Γ₀` is `supp v`. If `J` is an ideal of `R` +with `h : J ⊆ supp v` then the induced valuation +on R / J = `ideal.quotient J` is `on_quot v h`. + +-/ + +namespace valuation + +variables {R Γ₀ : Type*} [comm_ring R] [linear_ordered_comm_monoid_with_zero Γ₀] +variables (v : valuation R Γ₀) + +/-- If `hJ : J ⊆ supp v` then `on_quot_val hJ` is the induced function on R/J as a function. +Note: it's just the function; the valuation is `on_quot hJ`. -/ +def on_quot_val {J : ideal R} (hJ : J ≤ supp v) : + R ⧸ J → Γ₀ := +λ q, quotient.lift_on' q v $ λ a b h, +calc v a = v (b + -(-a + b)) : by simp + ... = v b : + v.map_add_supp b $ (ideal.neg_mem_iff _).2 $ hJ $ quotient_add_group.left_rel_apply.mp h + +/-- The extension of valuation v on R to valuation on R/J if J ⊆ supp v -/ +def on_quot {J : ideal R} (hJ : J ≤ supp v) : + valuation (R ⧸ J) Γ₀ := +{ to_fun := v.on_quot_val hJ, + map_zero' := v.map_zero, + map_one' := v.map_one, + map_mul' := λ xbar ybar, quotient.ind₂' v.map_mul xbar ybar, + map_add_le_max' := λ xbar ybar, quotient.ind₂' v.map_add xbar ybar } + +@[simp] lemma on_quot_comap_eq {J : ideal R} (hJ : J ≤ supp v) : + (v.on_quot hJ).comap (ideal.quotient.mk J) = v := +ext $ λ r, rfl + +lemma self_le_supp_comap (J : ideal R) (v : valuation (R ⧸ J) Γ₀) : + J ≤ (v.comap (ideal.quotient.mk J)).supp := +by { rw [comap_supp, ← ideal.map_le_iff_le_comap], simp } + +@[simp] lemma comap_on_quot_eq (J : ideal R) (v : valuation (R ⧸ J) Γ₀) : + (v.comap (ideal.quotient.mk J)).on_quot (v.self_le_supp_comap J) = v := +ext $ by { rintro ⟨x⟩, refl } + +/-- The quotient valuation on R/J has support supp(v)/J if J ⊆ supp v. -/ +lemma supp_quot {J : ideal R} (hJ : J ≤ supp v) : + supp (v.on_quot hJ) = (supp v).map (ideal.quotient.mk J) := +begin + apply le_antisymm, + { rintro ⟨x⟩ hx, + apply ideal.subset_span, + exact ⟨x, hx, rfl⟩ }, + { rw ideal.map_le_iff_le_comap, + intros x hx, exact hx } +end + +lemma supp_quot_supp : supp (v.on_quot le_rfl) = 0 := +by { rw supp_quot, exact ideal.map_quotient_self _ } + +end valuation + +namespace add_valuation + +variables {R Γ₀ : Type*} +variables [comm_ring R] [linear_ordered_add_comm_monoid_with_top Γ₀] +variables (v : add_valuation R Γ₀) + +local attribute [reducible] add_valuation + +/-- If `hJ : J ⊆ supp v` then `on_quot_val hJ` is the induced function on R/J as a function. +Note: it's just the function; the valuation is `on_quot hJ`. -/ +def on_quot_val {J : ideal R} (hJ : J ≤ supp v) : (R ⧸ J) → Γ₀ := v.on_quot_val hJ + +/-- The extension of valuation v on R to valuation on R/J if J ⊆ supp v -/ +def on_quot {J : ideal R} (hJ : J ≤ supp v) : + add_valuation (R ⧸ J) Γ₀ := +v.on_quot hJ + +@[simp] lemma on_quot_comap_eq {J : ideal R} (hJ : J ≤ supp v) : + (v.on_quot hJ).comap (ideal.quotient.mk J) = v := +v.on_quot_comap_eq hJ + +lemma comap_supp {S : Type*} [comm_ring S] (f : S →+* R) : + supp (v.comap f) = ideal.comap f v.supp := +v.comap_supp f + +lemma self_le_supp_comap (J : ideal R) (v : add_valuation (R ⧸ J) Γ₀) : + J ≤ (v.comap (ideal.quotient.mk J)).supp := +v.self_le_supp_comap J + +@[simp] lemma comap_on_quot_eq (J : ideal R) (v : add_valuation (R ⧸ J) Γ₀) : + (v.comap (ideal.quotient.mk J)).on_quot (v.self_le_supp_comap J) = v := +v.comap_on_quot_eq J + +/-- The quotient valuation on R/J has support supp(v)/J if J ⊆ supp v. -/ +lemma supp_quot {J : ideal R} (hJ : J ≤ supp v) : + supp (v.on_quot hJ) = (supp v).map (ideal.quotient.mk J) := +v.supp_quot hJ + +lemma supp_quot_supp : supp (v.on_quot le_rfl) = 0 := +v.supp_quot_supp + +end add_valuation From f24cc2891c0e328f0ee8c57387103aa462c44b5e Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sat, 11 Mar 2023 18:20:22 +0000 Subject: [PATCH 0614/1291] chore(data/finset/locally_finite): lemmas about open intervals (#18533) We had `cons` lemmas for the other four cases, but not these two. The new lemmas golf a proof a little. Also adds some docstrings to makes these lemmas easier to find. Forward-ported in https://github.com/leanprover-community/mathlib4/pull/2812 --- src/data/finset/locally_finite.lean | 22 ++++++++++++++++------ 1 file changed, 16 insertions(+), 6 deletions(-) diff --git a/src/data/finset/locally_finite.lean b/src/data/finset/locally_finite.lean index c495c94620e52..7366deb5165fa 100644 --- a/src/data/finset/locally_finite.lean +++ b/src/data/finset/locally_finite.lean @@ -332,12 +332,22 @@ end decidable_eq -- Those lemmas are purposefully the other way around +/-- `finset.cons` version of `finset.Ico_insert_right`. -/ lemma Icc_eq_cons_Ico (h : a ≤ b) : Icc a b = (Ico a b).cons b right_not_mem_Ico := by { classical, rw [cons_eq_insert, Ico_insert_right h] } +/-- `finset.cons` version of `finset.Ioc_insert_left`. -/ lemma Icc_eq_cons_Ioc (h : a ≤ b) : Icc a b = (Ioc a b).cons a left_not_mem_Ioc := by { classical, rw [cons_eq_insert, Ioc_insert_left h] } +/-- `finset.cons` version of `finset.Ioo_insert_right`. -/ +lemma Ioc_eq_cons_Ioo (h : a < b) : Ioc a b = (Ioo a b).cons b right_not_mem_Ioo := +by { classical, rw [cons_eq_insert, Ioo_insert_right h], } + +/-- `finset.cons` version of `finset.Ioo_insert_left`. -/ +lemma Ico_eq_cons_Ioo (h : a < b) : Ico a b = (Ioo a b).cons a left_not_mem_Ioo := +by { classical, rw [cons_eq_insert, Ioo_insert_left h] } + lemma Ico_filter_le_left {a b : α} [decidable_pred (≤ a)] (hab : a < b) : (Ico a b).filter (λ x, x ≤ a) = {a} := begin @@ -350,7 +360,7 @@ lemma card_Ico_eq_card_Icc_sub_one (a b : α) : (Ico a b).card = (Icc a b).card begin classical, by_cases h : a ≤ b, - { rw [←Ico_insert_right h, card_insert_of_not_mem right_not_mem_Ico], + { rw [Icc_eq_cons_Ico h, card_cons], exact (nat.add_sub_cancel _ _).symm }, { rw [Ico_eq_empty (λ h', h h'.le), Icc_eq_empty h, card_empty, zero_tsub] } end @@ -361,12 +371,10 @@ lemma card_Ioc_eq_card_Icc_sub_one (a b : α) : (Ioc a b).card = (Icc a b).card lemma card_Ioo_eq_card_Ico_sub_one (a b : α) : (Ioo a b).card = (Ico a b).card - 1 := begin classical, - by_cases h : a ≤ b, - { obtain rfl | h' := h.eq_or_lt, - { rw [Ioo_self, Ico_self, card_empty] }, - rw [←Ioo_insert_left h', card_insert_of_not_mem left_not_mem_Ioo], + by_cases h : a < b, + { rw [Ico_eq_cons_Ioo h, card_cons], exact (nat.add_sub_cancel _ _).symm }, - { rw [Ioo_eq_empty (λ h', h h'.le), Ico_eq_empty (λ h', h h'.le), card_empty, zero_tsub] } + { rw [Ioo_eq_empty h, Ico_eq_empty h, card_empty, zero_tsub] } end lemma card_Ioo_eq_card_Ioc_sub_one (a b : α) : (Ioo a b).card = (Ioc a b).card - 1 := @@ -391,6 +399,7 @@ by { ext, simp_rw [finset.mem_insert, mem_Ici, mem_Ioi, le_iff_lt_or_eq, or_comm @[simp] lemma not_mem_Ioi_self {b : α} : b ∉ Ioi b := λ h, lt_irrefl _ (mem_Ioi.1 h) -- Purposefully written the other way around +/-- `finset.cons` version of `finset.Ioi_insert`. -/ lemma Ici_eq_cons_Ioi (a : α) : Ici a = (Ioi a).cons a not_mem_Ioi_self := by { classical, rw [cons_eq_insert, Ioi_insert] } @@ -410,6 +419,7 @@ by { ext, simp_rw [finset.mem_insert, mem_Iic, mem_Iio, le_iff_lt_or_eq, or_comm @[simp] lemma not_mem_Iio_self {b : α} : b ∉ Iio b := λ h, lt_irrefl _ (mem_Iio.1 h) -- Purposefully written the other way around +/-- `finset.cons` version of `finset.Iio_insert`. -/ lemma Iic_eq_cons_Iio (b : α) : Iic b = (Iio b).cons b not_mem_Iio_self := by { classical, rw [cons_eq_insert, Iio_insert] } From c985ae9840e06836a71db38de372f20acb49b790 Mon Sep 17 00:00:00 2001 From: Eric Rodriguez Date: Sun, 12 Mar 2023 00:23:44 +0000 Subject: [PATCH 0615/1291] chore(topology/order/basic): generalise frontier_Icc (#18571) --- src/topology/order/basic.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/topology/order/basic.lean b/src/topology/order/basic.lean index 51e9d3489e4f6..733fa1a9de319 100644 --- a/src/topology/order/basic.lean +++ b/src/topology/order/basic.lean @@ -2132,9 +2132,9 @@ by simp [frontier, closure_Iio' ha, Iic_diff_Iio, Icc_self] lemma frontier_Iio [no_min_order α] {a : α} : frontier (Iio a) = {a} := frontier_Iio' nonempty_Iio -@[simp] lemma frontier_Icc [no_min_order α] [no_max_order α] {a b : α} (h : a < b) : +@[simp] lemma frontier_Icc [no_min_order α] [no_max_order α] {a b : α} (h : a ≤ b) : frontier (Icc a b) = {a, b} := -by simp [frontier, le_of_lt h, Icc_diff_Ioo_same] +by simp [frontier, h, Icc_diff_Ioo_same] @[simp] lemma frontier_Ioo {a b : α} (h : a < b) : frontier (Ioo a b) = {a, b} := by rw [frontier, closure_Ioo h.ne, interior_Ioo, Icc_diff_Ioo_same h.le] From 2af0836443b4cfb5feda0df0051acdb398304931 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sun, 12 Mar 2023 20:39:34 +0000 Subject: [PATCH 0616/1291] refactor(number_theory/zsqrtd): replace `zsqrtd.conj` with `star` (#18572) This allows more existing lemmas to be used; notably, `unitary (zqsrt d)` becomes something we can talk about. --- src/number_theory/pell_matiyasevic.lean | 16 ++++--- src/number_theory/zsqrtd/basic.lean | 54 +++++++--------------- src/number_theory/zsqrtd/gaussian_int.lean | 12 +++-- 3 files changed, 36 insertions(+), 46 deletions(-) diff --git a/src/number_theory/pell_matiyasevic.lean b/src/number_theory/pell_matiyasevic.lean index c36aeff75e744..014171f23684e 100644 --- a/src/number_theory/pell_matiyasevic.lean +++ b/src/number_theory/pell_matiyasevic.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ +import algebra.star.unitary import data.nat.modeq import number_theory.zsqrtd.basic @@ -121,15 +122,18 @@ theorem is_pell_nat {x y : ℕ} : is_pell ⟨x, y⟩ ↔ x*x - d*y*y = 1 := λh, show ((x*x : ℕ) - (d*y*y:ℕ) : ℤ) = 1, by rw [← int.coe_nat_sub $ le_of_lt $ nat.lt_of_sub_eq_succ h, h]; refl⟩ -theorem is_pell_norm : Π {b : ℤ√d}, is_pell b ↔ b * b.conj = 1 +theorem is_pell_norm : Π {b : ℤ√d}, is_pell b ↔ b * star b = 1 | ⟨x, y⟩ := by simp [zsqrtd.ext, is_pell, mul_comm]; ring_nf +theorem is_pell_iff_mem_unitary : Π {b : ℤ√d}, is_pell b ↔ b ∈ unitary ℤ√d +| ⟨x, y⟩ := by rw [unitary.mem_iff, is_pell_norm, mul_comm (star _), and_self] + theorem is_pell_mul {b c : ℤ√d} (hb : is_pell b) (hc : is_pell c) : is_pell (b * c) := is_pell_norm.2 (by simp [mul_comm, mul_left_comm, - zsqrtd.conj_mul, pell.is_pell_norm.1 hb, pell.is_pell_norm.1 hc]) + star_mul, pell.is_pell_norm.1 hb, pell.is_pell_norm.1 hc]) -theorem is_pell_conj : ∀ {b : ℤ√d}, is_pell b ↔ is_pell b.conj | ⟨x, y⟩ := -by simp [is_pell, zsqrtd.conj] +theorem is_pell_star : ∀ {b : ℤ√d}, is_pell b ↔ is_pell (star b) | ⟨x, y⟩ := +by simp [is_pell, zsqrtd.star_mk] @[simp] theorem pell_zd_succ (n : ℕ) : pell_zd (n+1) = pell_zd n * ⟨a, 1⟩ := by simp [zsqrtd.ext] @@ -187,7 +191,7 @@ lemma eq_pell_lem : ∀n (b:ℤ√d), 1 ≤ b → is_pell b → b ≤ pell_zd n if ha : (⟨↑a, 1⟩ : ℤ√d) ≤ b then let ⟨m, e⟩ := eq_pell_lem n (b * ⟨a, -1⟩) (by rw ← a1m; exact mul_le_mul_of_nonneg_right ha am1p) - (is_pell_mul hp (is_pell_conj.1 is_pell_one)) + (is_pell_mul hp (is_pell_star.1 is_pell_one)) (by have t := mul_le_mul_of_nonneg_right h am1p; rwa [pell_zd_succ, mul_assoc, a1m, mul_one] at t) in ⟨m+1, by rw [show b = b * ⟨a, -1⟩ * ⟨a, 1⟩, by rw [mul_assoc, eq.trans (mul_comm _ _) a1m]; @@ -245,7 +249,7 @@ by injection (pell_zd_add _ m n) with _ h; repeat {rw ← int.coe_nat_add at h <|> rw ← int.coe_nat_mul at h}; exact int.coe_nat_inj h -theorem pell_zd_sub {m n} (h : n ≤ m) : pell_zd (m - n) = pell_zd m * (pell_zd n).conj := +theorem pell_zd_sub {m n} (h : n ≤ m) : pell_zd (m - n) = pell_zd m * star (pell_zd n) := let t := pell_zd_add n (m - n) in by rw [add_tsub_cancel_of_le h] at t; rw [t, mul_comm (pell_zd _ n) _, mul_assoc, (is_pell_norm _).1 (is_pell_pell_zd _ _), mul_one] diff --git a/src/number_theory/zsqrtd/basic.lean b/src/number_theory/zsqrtd/basic.lean index 4e29530bb4f37..3047f61fef974 100644 --- a/src/number_theory/zsqrtd/basic.lean +++ b/src/number_theory/zsqrtd/basic.lean @@ -122,33 +122,16 @@ instance : ring ℤ√d := by apply_instance instance : distrib ℤ√d := by apply_instance /-- Conjugation in `ℤ√d`. The conjugate of `a + b √d` is `a - b √d`. -/ -def conj (z : ℤ√d) : ℤ√d := ⟨z.1, -z.2⟩ -@[simp] lemma conj_re (z : ℤ√d) : (conj z).re = z.re := rfl -@[simp] lemma conj_im (z : ℤ√d) : (conj z).im = -z.im := rfl +instance : has_star ℤ√d := +{ star := λ z, ⟨z.1, -z.2⟩ } +@[simp] lemma star_mk (x y : ℤ) : star (⟨x, y⟩ : ℤ√d) = ⟨x, -y⟩ := rfl +@[simp] lemma star_re (z : ℤ√d) : (star z).re = z.re := rfl +@[simp] lemma star_im (z : ℤ√d) : (star z).im = -z.im := rfl -/-- `conj` as an `add_monoid_hom`. -/ -def conj_hom : ℤ√d →+ ℤ√d := -{ to_fun := conj, - map_add' := λ ⟨a, ai⟩ ⟨b, bi⟩, ext.mpr ⟨rfl, neg_add _ _⟩, - map_zero' := ext.mpr ⟨rfl, neg_zero⟩ } - -@[simp] lemma conj_zero : conj (0 : ℤ√d) = 0 := -conj_hom.map_zero - -@[simp] lemma conj_one : conj (1 : ℤ√d) = 1 := -by simp only [zsqrtd.ext, zsqrtd.conj_re, zsqrtd.conj_im, zsqrtd.one_im, neg_zero, eq_self_iff_true, - and_self] - -@[simp] lemma conj_neg (x : ℤ√d) : (-x).conj = -x.conj := rfl - -@[simp] lemma conj_add (x y : ℤ√d) : (x + y).conj = x.conj + y.conj := -conj_hom.map_add x y - -@[simp] lemma conj_sub (x y : ℤ√d) : (x - y).conj = x.conj - y.conj := -conj_hom.map_sub x y - -@[simp] lemma conj_conj {d : ℤ} (x : ℤ√d) : x.conj.conj = x := -by simp only [ext, true_and, conj_re, eq_self_iff_true, neg_neg, conj_im] +instance : star_ring ℤ√d := +{ star_involutive := λ x, ext.mpr ⟨rfl, neg_neg _⟩, + star_mul := λ a b, ext.mpr ⟨by simp; ring, by simp; ring⟩, + star_add := λ a b, ext.mpr ⟨rfl, neg_add _ _⟩ } instance : nontrivial ℤ√d := ⟨⟨0, 1, dec_trivial⟩⟩ @@ -188,12 +171,9 @@ by simp [ext] theorem decompose {x y : ℤ} : (⟨x, y⟩ : ℤ√d) = x + sqrtd * y := by simp [ext] -theorem mul_conj {x y : ℤ} : (⟨x, y⟩ * conj ⟨x, y⟩ : ℤ√d) = x * x - d * y * y := +theorem mul_star {x y : ℤ} : (⟨x, y⟩ * star ⟨x, y⟩ : ℤ√d) = x * x - d * y * y := by simp [ext, sub_eq_add_neg, mul_comm] -theorem conj_mul {a b : ℤ√d} : conj (a * b) = conj a * conj b := -by { simp [ext], ring } - protected lemma coe_int_add (m n : ℤ) : (↑(m + n) : ℤ√d) = ↑m + ↑n := (int.cast_ring_hom _).map_add _ _ protected lemma coe_int_sub (m n : ℤ) : (↑(m - n) : ℤ√d) = ↑m - ↑n := @@ -376,15 +356,15 @@ def norm_monoid_hom : ℤ√d →* ℤ := map_mul' := norm_mul, map_one' := norm_one } -lemma norm_eq_mul_conj (n : ℤ√d) : (norm n : ℤ√d) = n * n.conj := -by cases n; simp [norm, conj, zsqrtd.ext, mul_comm, sub_eq_add_neg] +lemma norm_eq_mul_conj (n : ℤ√d) : (norm n : ℤ√d) = n * star n := +by cases n; simp [norm, star, zsqrtd.ext, mul_comm, sub_eq_add_neg] @[simp] lemma norm_neg (x : ℤ√d) : (-x).norm = x.norm := -coe_int_inj $ by simp only [norm_eq_mul_conj, conj_neg, neg_mul, +coe_int_inj $ by simp only [norm_eq_mul_conj, star_neg, neg_mul, mul_neg, neg_neg] -@[simp] lemma norm_conj (x : ℤ√d) : x.conj.norm = x.norm := -coe_int_inj $ by simp only [norm_eq_mul_conj, conj_conj, mul_comm] +@[simp] lemma norm_conj (x : ℤ√d) : (star x).norm = x.norm := +coe_int_inj $ by simp only [norm_eq_mul_conj, star_star, mul_comm] lemma norm_nonneg (hd : d ≤ 0) (n : ℤ√d) : 0 ≤ n.norm := add_nonneg (mul_self_nonneg _) @@ -394,10 +374,10 @@ add_nonneg (mul_self_nonneg _) lemma norm_eq_one_iff {x : ℤ√d} : x.norm.nat_abs = 1 ↔ is_unit x := ⟨λ h, is_unit_iff_dvd_one.2 $ (le_total 0 (norm x)).cases_on - (λ hx, show x ∣ 1, from ⟨x.conj, + (λ hx, show x ∣ 1, from ⟨star x, by rwa [← int.coe_nat_inj', int.nat_abs_of_nonneg hx, ← @int.cast_inj (ℤ√d) _ _, norm_eq_mul_conj, eq_comm] at h⟩) - (λ hx, show x ∣ 1, from ⟨- x.conj, + (λ hx, show x ∣ 1, from ⟨- star x, by rwa [← int.coe_nat_inj', int.of_nat_nat_abs_of_nonpos hx, ← @int.cast_inj (ℤ√d) _ _, int.cast_neg, norm_eq_mul_conj, neg_mul_eq_mul_neg, eq_comm] at h⟩), diff --git a/src/number_theory/zsqrtd/gaussian_int.lean b/src/number_theory/zsqrtd/gaussian_int.lean index 8aa9a9f782990..363cfc9dad6ca 100644 --- a/src/number_theory/zsqrtd/gaussian_int.lean +++ b/src/number_theory/zsqrtd/gaussian_int.lean @@ -36,6 +36,7 @@ and definitions about `zsqrtd` can easily be used. -/ open zsqrtd complex +open_locale complex_conjugate /-- The Gaussian integers, defined as `ℤ√(-1)`. -/ @[reducible] def gaussian_int : Type := zsqrtd (-1) @@ -75,6 +76,11 @@ by apply complex.ext; simp [to_complex_def] @[simp] lemma to_complex_zero : ((0 : ℤ[i]) : ℂ) = 0 := to_complex.map_zero @[simp] lemma to_complex_neg (x : ℤ[i]) : ((-x : ℤ[i]) : ℂ) = -x := to_complex.map_neg _ @[simp] lemma to_complex_sub (x y : ℤ[i]) : ((x - y : ℤ[i]) : ℂ) = x - y := to_complex.map_sub _ _ +@[simp] lemma to_complex_star (x : ℤ[i]) : ((star x : ℤ[i]) : ℂ) = conj (x : ℂ) := +begin + rw [to_complex_def₂, to_complex_def₂], + exact congr_arg2 _ rfl (int.cast_neg _), +end @[simp] lemma to_complex_inj {x y : ℤ[i]} : (x : ℂ) = y ↔ x = y := by cases x; cases y; simp [to_complex_def₂] @@ -108,11 +114,11 @@ lemma nat_abs_norm_eq (x : ℤ[i]) : x.norm.nat_abs = int.coe_nat_inj $ begin simp, simp [zsqrtd.norm] end instance : has_div ℤ[i] := -⟨λ x y, let n := (norm y : ℚ)⁻¹, c := y.conj in +⟨λ x y, let n := (norm y : ℚ)⁻¹, c := star y in ⟨round ((x * c).re * n : ℚ), round ((x * c).im * n : ℚ)⟩⟩ -lemma div_def (x y : ℤ[i]) : x / y = ⟨round ((x * conj y).re / norm y : ℚ), - round ((x * conj y).im / norm y : ℚ)⟩ := +lemma div_def (x y : ℤ[i]) : x / y = ⟨round ((x * star y).re / norm y : ℚ), + round ((x * star y).im / norm y : ℚ)⟩ := show zsqrtd.mk _ _ = _, by simp [div_eq_mul_inv] lemma to_complex_div_re (x y : ℤ[i]) : ((x / y : ℤ[i]) : ℂ).re = round ((x / y : ℂ).re) := From 1e3201306d4d9eb1fd54c60d7c4510ad5126f6f9 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sun, 12 Mar 2023 22:38:09 +0000 Subject: [PATCH 0617/1291] feat(analysis/normed_space/*exponential): more results about `star` (#18553) These are trivial consequences of existing results. It turns out we already had a proof about `exp _ x` belonging to `unitary X`; this weakens the conditions. --- src/algebra/star/self_adjoint.lean | 17 +++++++++++++++++ src/analysis/normed_space/exponential.lean | 12 ++++++++++++ .../normed_space/matrix_exponential.lean | 9 +++++++++ src/analysis/normed_space/star/exponential.lean | 12 +----------- src/analysis/normed_space/star/spectrum.lean | 5 +++-- 5 files changed, 42 insertions(+), 13 deletions(-) diff --git a/src/algebra/star/self_adjoint.lean b/src/algebra/star/self_adjoint.lean index e0dbbff87d05e..d0dc47e4e4988 100644 --- a/src/algebra/star/self_adjoint.lean +++ b/src/algebra/star/self_adjoint.lean @@ -34,6 +34,7 @@ We also define `is_star_normal R`, a `Prop` that states that an element `x` sati ## TODO +* Define `is_skew_adjoint` to match `is_self_adjoint`. * Define `λ z x, z * x * star z` (i.e. conjugation by `z`) as a monoid action of `R` on `R` (similar to the existing `conj_act` for groups), and then state the fact that `self_adjoint R` is invariant under it. @@ -399,6 +400,22 @@ end has_smul end skew_adjoint +/-- Scalar multiplication of a self-adjoint element by a skew-adjoint element produces a +skew-adjoint element. -/ +lemma is_self_adjoint.smul_mem_skew_adjoint [ring R] [add_comm_group A] [module R A] + [star_add_monoid R] [star_add_monoid A] [star_module R A] {r : R} + (hr : r ∈ skew_adjoint R) {a : A} (ha : is_self_adjoint a) : + r • a ∈ skew_adjoint A := +(star_smul _ _).trans $ (congr_arg2 _ hr ha).trans $ neg_smul _ _ + +/-- Scalar multiplication of a skew-adjoint element by a skew-adjoint element produces a +self-adjoint element. -/ +lemma is_self_adjoint_smul_of_mem_skew_adjoint [ring R] [add_comm_group A] [module R A] + [star_add_monoid R] [star_add_monoid A] [star_module R A] {r : R} + (hr : r ∈ skew_adjoint R) {a : A} (ha : a ∈ skew_adjoint A) : + is_self_adjoint (r • a) := +(star_smul _ _).trans $ (congr_arg2 _ hr ha).trans $ neg_smul_neg _ _ + instance is_star_normal_zero [semiring R] [star_ring R] : is_star_normal (0 : R) := ⟨by simp only [star_comm_self, star_zero]⟩ diff --git a/src/analysis/normed_space/exponential.lean b/src/analysis/normed_space/exponential.lean index 31473c07460cc..ed56dd6b68828 100644 --- a/src/analysis/normed_space/exponential.lean +++ b/src/analysis/normed_space/exponential.lean @@ -125,6 +125,11 @@ by simp_rw [exp_eq_tsum, ←star_pow, ←star_inv_nat_cast_smul, ←tsum_star] variables (𝕂) +lemma is_self_adjoint.exp [t2_space 𝔸] [star_ring 𝔸] [has_continuous_star 𝔸] {x : 𝔸} + (h : is_self_adjoint x) : + is_self_adjoint (exp 𝕂 x) := +(star_exp x).trans $ h.symm ▸ rfl + lemma commute.exp_right [t2_space 𝔸] {x y : 𝔸} (h : commute x y) : commute x (exp 𝕂 y) := begin rw exp_eq_tsum, @@ -447,6 +452,13 @@ begin exact ring.inverse_invertible _, end +lemma exp_mem_unitary_of_mem_skew_adjoint [star_ring 𝔸] [has_continuous_star 𝔸] {x : 𝔸} + (h : x ∈ skew_adjoint 𝔸) : + exp 𝕂 x ∈ unitary 𝔸 := +by rw [unitary.mem_iff, star_exp, skew_adjoint.mem_iff.mp h, + ←exp_add_of_commute (commute.refl x).neg_left, ←exp_add_of_commute (commute.refl x).neg_right, + add_left_neg, add_right_neg, exp_zero, and_self] + end /-- In a Banach-algebra `𝔸` over `𝕂 = ℝ` or `𝕂 = ℂ`, if a family of elements `f i` mutually diff --git a/src/analysis/normed_space/matrix_exponential.lean b/src/analysis/normed_space/matrix_exponential.lean index 423ad6dfbd9ea..529ec5ab24174 100644 --- a/src/analysis/normed_space/matrix_exponential.lean +++ b/src/analysis/normed_space/matrix_exponential.lean @@ -7,6 +7,8 @@ Authors: Eric Wieser import analysis.normed_space.exponential import analysis.matrix import linear_algebra.matrix.zpow +import linear_algebra.matrix.hermitian +import linear_algebra.matrix.symmetric import topology.uniform_space.matrix /-! @@ -118,6 +120,10 @@ lemma exp_conj_transpose [star_ring 𝔸] [has_continuous_star 𝔸] (A : matrix exp 𝕂 Aᴴ = (exp 𝕂 A)ᴴ := (star_exp A).symm +lemma is_hermitian.exp [star_ring 𝔸] [has_continuous_star 𝔸] {A : matrix m m 𝔸} + (h : A.is_hermitian) : (exp 𝕂 A).is_hermitian := +(exp_conj_transpose _ _).symm.trans $ congr_arg _ h + end ring section comm_ring @@ -127,6 +133,9 @@ variables [fintype m] [decidable_eq m] [field 𝕂] lemma exp_transpose (A : matrix m m 𝔸) : exp 𝕂 Aᵀ = (exp 𝕂 A)ᵀ := by simp_rw [exp_eq_tsum, transpose_tsum, transpose_smul, transpose_pow] +lemma is_symm.exp {A : matrix m m 𝔸} (h : A.is_symm) : (exp 𝕂 A).is_symm := +(exp_transpose _ _).symm.trans $ congr_arg _ h + end comm_ring end topological diff --git a/src/analysis/normed_space/star/exponential.lean b/src/analysis/normed_space/star/exponential.lean index 1cd4c9045f96b..36b05dd9cf7b1 100644 --- a/src/analysis/normed_space/star/exponential.lean +++ b/src/analysis/normed_space/star/exponential.lean @@ -26,21 +26,11 @@ variables {A : Type*} open complex -lemma is_self_adjoint.exp_i_smul_unitary {a : A} (ha : is_self_adjoint a) : - exp ℂ (I • a) ∈ unitary A := -begin - rw [unitary.mem_iff, star_exp], - simp only [star_smul, is_R_or_C.star_def, self_adjoint.mem_iff.mp ha, conj_I, neg_smul], - rw ←@exp_add_of_commute ℂ A _ _ _ _ _ _ ((commute.refl (I • a)).neg_left), - rw ←@exp_add_of_commute ℂ A _ _ _ _ _ _ ((commute.refl (I • a)).neg_right), - simpa only [add_right_neg, add_left_neg, and_self] using (exp_zero : exp ℂ (0 : A) = 1), -end - /-- The map from the selfadjoint real subspace to the unitary group. This map only makes sense over ℂ. -/ @[simps] noncomputable def self_adjoint.exp_unitary (a : self_adjoint A) : unitary A := -⟨exp ℂ (I • a), a.prop.exp_i_smul_unitary⟩ +⟨exp ℂ (I • a), exp_mem_unitary_of_mem_skew_adjoint _ (a.prop.smul_mem_skew_adjoint conj_I)⟩ open self_adjoint diff --git a/src/analysis/normed_space/star/spectrum.lean b/src/analysis/normed_space/star/spectrum.lean index 1922301bca137..8994342eb2cb8 100644 --- a/src/analysis/normed_space/star/spectrum.lean +++ b/src/analysis/normed_space/star/spectrum.lean @@ -10,7 +10,7 @@ import analysis.special_functions.exponential import algebra.star.star_alg_hom /-! # Spectral properties in C⋆-algebras -In this file, we establish various propreties related to the spectrum of elements in C⋆-algebras. +In this file, we establish various properties related to the spectrum of elements in C⋆-algebras. -/ local postfix `⋆`:std.prec.max_plus := star @@ -90,6 +90,7 @@ end theorem is_self_adjoint.mem_spectrum_eq_re [star_module ℂ A] {a : A} (ha : is_self_adjoint a) {z : ℂ} (hz : z ∈ spectrum ℂ a) : z = z.re := begin + have hu := exp_mem_unitary_of_mem_skew_adjoint ℂ (ha.smul_mem_skew_adjoint conj_I), let Iu := units.mk0 I I_ne_zero, have : exp ℂ (I • z) ∈ spectrum ℂ (exp ℂ (I • a)), by simpa only [units.smul_def, units.coe_mk0] @@ -97,7 +98,7 @@ begin exact complex.ext (of_real_re _) (by simpa only [←complex.exp_eq_exp_ℂ, mem_sphere_zero_iff_norm, norm_eq_abs, abs_exp, real.exp_eq_one_iff, smul_eq_mul, I_mul, neg_eq_zero] - using spectrum.subset_circle_of_unitary ha.exp_i_smul_unitary this), + using spectrum.subset_circle_of_unitary hu this), end /-- Any element of the spectrum of a selfadjoint is real. -/ From 795b501869b9fa7aa716d5fdadd00c03f983a605 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 13 Mar 2023 05:01:01 +0000 Subject: [PATCH 0618/1291] chore(number_theory/pell_matiyasevic): generalize `is_pell` to arbitary `d` (#18573) Not much else in the file generalizes, but we may as well generalize the definition. --- src/number_theory/pell_matiyasevic.lean | 44 ++++++++++++++----------- 1 file changed, 24 insertions(+), 20 deletions(-) diff --git a/src/number_theory/pell_matiyasevic.lean b/src/number_theory/pell_matiyasevic.lean index 014171f23684e..76830c1fc9049 100644 --- a/src/number_theory/pell_matiyasevic.lean +++ b/src/number_theory/pell_matiyasevic.lean @@ -52,6 +52,27 @@ Pell's equation, Matiyasevic's theorem, Hilbert's tenth problem namespace pell open nat +section +variables {d : ℤ} + +/-- The property of being a solution to the Pell equation, expressed + as a property of elements of `ℤ√d`. -/ +def is_pell : ℤ√d → Prop | ⟨x, y⟩ := x*x - d*y*y = 1 + +theorem is_pell_norm : Π {b : ℤ√d}, is_pell b ↔ b * star b = 1 +| ⟨x, y⟩ := by simp [zsqrtd.ext, is_pell, mul_comm]; ring_nf + +theorem is_pell_iff_mem_unitary : Π {b : ℤ√d}, is_pell b ↔ b ∈ unitary ℤ√d +| ⟨x, y⟩ := by rw [unitary.mem_iff, is_pell_norm, mul_comm (star _), and_self] + +theorem is_pell_mul {b c : ℤ√d} (hb : is_pell b) (hc : is_pell c) : is_pell (b * c) := +is_pell_norm.2 (by simp [mul_comm, mul_left_comm, star_mul, is_pell_norm.1 hb, is_pell_norm.1 hc]) + +theorem is_pell_star : ∀ {b : ℤ√d}, is_pell b ↔ is_pell (star b) | ⟨x, y⟩ := +by simp [is_pell, zsqrtd.star_mk] + +end + section parameters {a : ℕ} (a1 : 1 < a) @@ -112,33 +133,16 @@ def pell_zd (n : ℕ) : ℤ√d := ⟨xn n, yn n⟩ @[simp] theorem pell_zd_re (n : ℕ) : (pell_zd n).re = xn n := rfl @[simp] theorem pell_zd_im (n : ℕ) : (pell_zd n).im = yn n := rfl -/-- The property of being a solution to the Pell equation, expressed - as a property of elements of `ℤ√d`. -/ -def is_pell : ℤ√d → Prop | ⟨x, y⟩ := x*x - d*y*y = 1 - -theorem is_pell_nat {x y : ℕ} : is_pell ⟨x, y⟩ ↔ x*x - d*y*y = 1 := +theorem is_pell_nat {x y : ℕ} : is_pell (⟨x, y⟩ : ℤ√d) ↔ x*x - d*y*y = 1 := ⟨λh, int.coe_nat_inj (by rw int.coe_nat_sub (int.le_of_coe_nat_le_coe_nat $ int.le.intro_sub h); exact h), λh, show ((x*x : ℕ) - (d*y*y:ℕ) : ℤ) = 1, by rw [← int.coe_nat_sub $ le_of_lt $ nat.lt_of_sub_eq_succ h, h]; refl⟩ -theorem is_pell_norm : Π {b : ℤ√d}, is_pell b ↔ b * star b = 1 -| ⟨x, y⟩ := by simp [zsqrtd.ext, is_pell, mul_comm]; ring_nf - -theorem is_pell_iff_mem_unitary : Π {b : ℤ√d}, is_pell b ↔ b ∈ unitary ℤ√d -| ⟨x, y⟩ := by rw [unitary.mem_iff, is_pell_norm, mul_comm (star _), and_self] - -theorem is_pell_mul {b c : ℤ√d} (hb : is_pell b) (hc : is_pell c) : is_pell (b * c) := -is_pell_norm.2 (by simp [mul_comm, mul_left_comm, - star_mul, pell.is_pell_norm.1 hb, pell.is_pell_norm.1 hc]) - -theorem is_pell_star : ∀ {b : ℤ√d}, is_pell b ↔ is_pell (star b) | ⟨x, y⟩ := -by simp [is_pell, zsqrtd.star_mk] - @[simp] theorem pell_zd_succ (n : ℕ) : pell_zd (n+1) = pell_zd n * ⟨a, 1⟩ := by simp [zsqrtd.ext] -theorem is_pell_one : is_pell ⟨a, 1⟩ := +theorem is_pell_one : is_pell (⟨a, 1⟩ : ℤ√d) := show az*az-d*1*1=1, by simp [dz_val]; ring theorem is_pell_pell_zd : ∀ (n : ℕ), is_pell (pell_zd n) @@ -252,7 +256,7 @@ by injection (pell_zd_add _ m n) with _ h; theorem pell_zd_sub {m n} (h : n ≤ m) : pell_zd (m - n) = pell_zd m * star (pell_zd n) := let t := pell_zd_add n (m - n) in by rw [add_tsub_cancel_of_le h] at t; - rw [t, mul_comm (pell_zd _ n) _, mul_assoc, (is_pell_norm _).1 (is_pell_pell_zd _ _), mul_one] + rw [t, mul_comm (pell_zd _ n) _, mul_assoc, is_pell_norm.1 (is_pell_pell_zd _ _), mul_one] theorem xz_sub {m n} (h : n ≤ m) : xz (m - n) = xz m * xz n - d * yz m * yz n := by { rw [sub_eq_add_neg, ←mul_neg], exact congr_arg zsqrtd.re (pell_zd_sub a1 h) } From 69c6a5a12d8a2b159f20933e60115a4f2de62b58 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Mon, 13 Mar 2023 06:41:40 +0000 Subject: [PATCH 0619/1291] chore(*): add mathlib4 synchronization comments (#18566) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.algebra.subalgebra.basic` * `algebra.algebra.subalgebra.tower` * `algebra.free_non_unital_non_assoc_algebra` * `algebra.monoid_algebra.basic` * `algebra.monoid_algebra.support` * `algebra.polynomial.big_operators` * `algebra.star.subalgebra` * `analysis.normed.group.ball_sphere` * `analysis.normed.group.basic` * `analysis.normed.group.completion` * `analysis.normed.group.hom` * `analysis.normed.group.hom_completion` * `analysis.normed.group.infinite_sum` * `analysis.normed.group.seminorm` * `analysis.normed_space.indicator_function` * `analysis.subadditive` * `category_theory.abelian.images` * `category_theory.abelian.non_preadditive` * `category_theory.adjunction.comma` * `category_theory.concrete_category.reflects_isomorphisms` * `category_theory.is_connected` * `category_theory.limits.bicones` * `category_theory.limits.comma` * `category_theory.limits.cone_category` * `category_theory.limits.constructions.equalizers` * `category_theory.limits.constructions.finite_products_of_binary_products` * `category_theory.limits.constructions.limits_of_products_and_equalizers` * `category_theory.limits.essentially_small` * `category_theory.limits.kan_extension` * `category_theory.limits.lattice` * `category_theory.limits.mono_coprod` * `category_theory.limits.opposites` * `category_theory.limits.preserves.filtered` * `category_theory.limits.preserves.shapes.biproducts` * `category_theory.limits.preserves.shapes.zero` * `category_theory.limits.shapes.biproducts` * `category_theory.limits.shapes.finite_products` * `category_theory.limits.shapes.functor_category` * `category_theory.limits.shapes.kernels` * `category_theory.limits.shapes.normal_mono.basic` * `category_theory.limits.shapes.normal_mono.equalizers` * `category_theory.limits.shapes.zero_morphisms` * `category_theory.limits.small_complete` * `category_theory.limits.types` * `category_theory.linear.basic` * `category_theory.linear.functor_category` * `category_theory.monoidal.End` * `category_theory.preadditive.additive_functor` * `category_theory.preadditive.basic` * `category_theory.preadditive.biproducts` * `category_theory.preadditive.functor_category` * `combinatorics.hindman` * `control.bifunctor` * `control.bitraversable.basic` * `data.complex.basic` * `data.nat.choose.cast` * `data.nat.choose.vandermonde` * `data.nat.factorial.cast` * `data.polynomial.basic` * `data.polynomial.cardinal` * `data.polynomial.coeff` * `data.polynomial.degree.definitions` * `data.polynomial.degree.lemmas` * `data.polynomial.degree.trailing_degree` * `data.polynomial.derivative` * `data.polynomial.erase_lead` * `data.polynomial.eval` * `data.polynomial.induction` * `data.polynomial.inductions` * `data.polynomial.monic` * `data.polynomial.monomial` * `data.polynomial.reverse` * `data.real.sqrt` * `data.tree` * `linear_algebra.affine_space.slope` * `logic.equiv.functor` * `number_theory.basic` * `order.jordan_holder` * `ring_theory.adjoin.basic` * `ring_theory.ideal.operations` * `ring_theory.ideal.quotient` * `ring_theory.localization.basic` * `ring_theory.nilpotent` * `ring_theory.polynomial.pochhammer` * `topology.algebra.affine` * `topology.algebra.localization` * `topology.algebra.ring.ideal` * `topology.fiber_bundle.trivialization` * `topology.instances.ennreal` * `topology.metric_space.completion` * `topology.metric_space.gluing` * `topology.metric_space.isometric_smul` * `topology.metric_space.isometry` * `topology.uniform_space.compare_reals` * `topology.unit_interval` --- src/algebra/algebra/subalgebra/basic.lean | 3 +++ src/algebra/algebra/subalgebra/tower.lean | 3 +++ src/algebra/free_non_unital_non_assoc_algebra.lean | 3 +++ src/algebra/monoid_algebra/basic.lean | 3 +++ src/algebra/monoid_algebra/support.lean | 3 +++ src/algebra/polynomial/big_operators.lean | 3 +++ src/algebra/star/subalgebra.lean | 3 +++ src/analysis/normed/group/ball_sphere.lean | 3 +++ src/analysis/normed/group/basic.lean | 3 +++ src/analysis/normed/group/completion.lean | 3 +++ src/analysis/normed/group/hom.lean | 3 +++ src/analysis/normed/group/hom_completion.lean | 3 +++ src/analysis/normed/group/infinite_sum.lean | 3 +++ src/analysis/normed/group/seminorm.lean | 3 +++ src/analysis/normed_space/indicator_function.lean | 3 +++ src/analysis/subadditive.lean | 3 +++ src/category_theory/abelian/images.lean | 3 +++ src/category_theory/abelian/non_preadditive.lean | 3 +++ src/category_theory/adjunction/comma.lean | 3 +++ .../concrete_category/reflects_isomorphisms.lean | 3 +++ src/category_theory/is_connected.lean | 3 +++ src/category_theory/limits/bicones.lean | 3 +++ src/category_theory/limits/comma.lean | 3 +++ src/category_theory/limits/cone_category.lean | 3 +++ src/category_theory/limits/constructions/equalizers.lean | 3 +++ .../constructions/finite_products_of_binary_products.lean | 3 +++ .../constructions/limits_of_products_and_equalizers.lean | 3 +++ src/category_theory/limits/essentially_small.lean | 3 +++ src/category_theory/limits/kan_extension.lean | 3 +++ src/category_theory/limits/lattice.lean | 3 +++ src/category_theory/limits/mono_coprod.lean | 3 +++ src/category_theory/limits/opposites.lean | 3 +++ src/category_theory/limits/preserves/filtered.lean | 3 +++ src/category_theory/limits/preserves/shapes/biproducts.lean | 3 +++ src/category_theory/limits/preserves/shapes/zero.lean | 3 +++ src/category_theory/limits/shapes/biproducts.lean | 3 +++ src/category_theory/limits/shapes/finite_products.lean | 3 +++ src/category_theory/limits/shapes/functor_category.lean | 3 +++ src/category_theory/limits/shapes/kernels.lean | 3 +++ src/category_theory/limits/shapes/normal_mono/basic.lean | 3 +++ src/category_theory/limits/shapes/normal_mono/equalizers.lean | 3 +++ src/category_theory/limits/shapes/zero_morphisms.lean | 3 +++ src/category_theory/limits/small_complete.lean | 3 +++ src/category_theory/limits/types.lean | 3 +++ src/category_theory/linear/basic.lean | 3 +++ src/category_theory/linear/functor_category.lean | 3 +++ src/category_theory/monoidal/End.lean | 3 +++ src/category_theory/preadditive/additive_functor.lean | 3 +++ src/category_theory/preadditive/basic.lean | 3 +++ src/category_theory/preadditive/biproducts.lean | 3 +++ src/category_theory/preadditive/functor_category.lean | 3 +++ src/combinatorics/hindman.lean | 3 +++ src/control/bifunctor.lean | 3 +++ src/control/bitraversable/basic.lean | 3 +++ src/data/complex/basic.lean | 3 +++ src/data/nat/choose/cast.lean | 3 +++ src/data/nat/choose/vandermonde.lean | 3 +++ src/data/nat/factorial/cast.lean | 3 +++ src/data/polynomial/basic.lean | 3 +++ src/data/polynomial/cardinal.lean | 3 +++ src/data/polynomial/coeff.lean | 3 +++ src/data/polynomial/degree/definitions.lean | 3 +++ src/data/polynomial/degree/lemmas.lean | 3 +++ src/data/polynomial/degree/trailing_degree.lean | 3 +++ src/data/polynomial/derivative.lean | 3 +++ src/data/polynomial/erase_lead.lean | 3 +++ src/data/polynomial/eval.lean | 3 +++ src/data/polynomial/induction.lean | 3 +++ src/data/polynomial/inductions.lean | 3 +++ src/data/polynomial/monic.lean | 3 +++ src/data/polynomial/monomial.lean | 3 +++ src/data/polynomial/reverse.lean | 3 +++ src/data/real/sqrt.lean | 3 +++ src/data/tree.lean | 3 +++ src/linear_algebra/affine_space/slope.lean | 3 +++ src/logic/equiv/functor.lean | 3 +++ src/number_theory/basic.lean | 3 +++ src/order/jordan_holder.lean | 3 +++ src/ring_theory/adjoin/basic.lean | 3 +++ src/ring_theory/ideal/operations.lean | 3 +++ src/ring_theory/ideal/quotient.lean | 3 +++ src/ring_theory/localization/basic.lean | 3 +++ src/ring_theory/nilpotent.lean | 3 +++ src/ring_theory/polynomial/pochhammer.lean | 3 +++ src/topology/algebra/affine.lean | 3 +++ src/topology/algebra/localization.lean | 3 +++ src/topology/algebra/ring/ideal.lean | 3 +++ src/topology/fiber_bundle/trivialization.lean | 3 +++ src/topology/instances/ennreal.lean | 3 +++ src/topology/metric_space/completion.lean | 3 +++ src/topology/metric_space/gluing.lean | 3 +++ src/topology/metric_space/isometric_smul.lean | 3 +++ src/topology/metric_space/isometry.lean | 3 +++ src/topology/uniform_space/compare_reals.lean | 3 +++ src/topology/unit_interval.lean | 3 +++ 95 files changed, 285 insertions(+) diff --git a/src/algebra/algebra/subalgebra/basic.lean b/src/algebra/algebra/subalgebra/basic.lean index e4b47764f34e5..b0ca846d8f4d1 100644 --- a/src/algebra/algebra/subalgebra/basic.lean +++ b/src/algebra/algebra/subalgebra/basic.lean @@ -11,6 +11,9 @@ import ring_theory.ideal.operations /-! # Subalgebras over Commutative Semiring +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define `subalgebra`s and the usual operations on them (`map`, `comap`). More lemmas about `adjoin` can be found in `ring_theory.adjoin`. diff --git a/src/algebra/algebra/subalgebra/tower.lean b/src/algebra/algebra/subalgebra/tower.lean index 8edd9a4b26b94..a66175054ee33 100644 --- a/src/algebra/algebra/subalgebra/tower.lean +++ b/src/algebra/algebra/subalgebra/tower.lean @@ -10,6 +10,9 @@ import algebra.algebra.tower /-! # Subalgebras in towers of algebras +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove facts about subalgebras in towers of algebra. An algebra tower A/S/R is expressed by having instances of `algebra A S`, diff --git a/src/algebra/free_non_unital_non_assoc_algebra.lean b/src/algebra/free_non_unital_non_assoc_algebra.lean index 3b1a15738d151..4bb1d33b6ebda 100644 --- a/src/algebra/free_non_unital_non_assoc_algebra.lean +++ b/src/algebra/free_non_unital_non_assoc_algebra.lean @@ -9,6 +9,9 @@ import algebra.monoid_algebra.basic /-! # Free algebras +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given a semiring `R` and a type `X`, we construct the free non-unital, non-associative algebra on `X` with coefficients in `R`, together with its universal property. The construction is valuable because it can be used to build free algebras with more structure, e.g., free Lie algebras. diff --git a/src/algebra/monoid_algebra/basic.lean b/src/algebra/monoid_algebra/basic.lean index 6df0621c49e15..d68da00ac395c 100644 --- a/src/algebra/monoid_algebra/basic.lean +++ b/src/algebra/monoid_algebra/basic.lean @@ -12,6 +12,9 @@ import linear_algebra.finsupp /-! # Monoid algebras +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + When the domain of a `finsupp` has a multiplicative or additive structure, we can define a convolution product. To mathematicians this structure is known as the "monoid algebra", i.e. the finite formal linear combinations over a given semiring of elements of the monoid. diff --git a/src/algebra/monoid_algebra/support.lean b/src/algebra/monoid_algebra/support.lean index cb6909ab07316..9d94e76b6fb2a 100644 --- a/src/algebra/monoid_algebra/support.lean +++ b/src/algebra/monoid_algebra/support.lean @@ -7,6 +7,9 @@ import algebra.monoid_algebra.basic /-! # Lemmas about the support of a finitely supported function + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ universes u₁ u₂ u₃ diff --git a/src/algebra/polynomial/big_operators.lean b/src/algebra/polynomial/big_operators.lean index f6fdd04502af9..27bbec5a49037 100644 --- a/src/algebra/polynomial/big_operators.lean +++ b/src/algebra/polynomial/big_operators.lean @@ -8,6 +8,9 @@ import data.polynomial.monic /-! # Lemmas for the interaction between polynomials and `∑` and `∏`. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Recall that `∑` and `∏` are notation for `finset.sum` and `finset.prod` respectively. ## Main results diff --git a/src/algebra/star/subalgebra.lean b/src/algebra/star/subalgebra.lean index b2c0793876a2f..7356d1122fe09 100644 --- a/src/algebra/star/subalgebra.lean +++ b/src/algebra/star/subalgebra.lean @@ -12,6 +12,9 @@ import ring_theory.adjoin.basic /-! # Star subalgebras +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A *-subalgebra is a subalgebra of a *-algebra which is closed under *. The centralizer of a *-closed set is a *-subalgebra. diff --git a/src/analysis/normed/group/ball_sphere.lean b/src/analysis/normed/group/ball_sphere.lean index 64d63bc2ecf10..6d90e2110fa4b 100644 --- a/src/analysis/normed/group/ball_sphere.lean +++ b/src/analysis/normed/group/ball_sphere.lean @@ -8,6 +8,9 @@ import analysis.normed.group.basic /-! # Negation on spheres and balls +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define `has_involutive_neg` instances for spheres, open balls, and closed balls in a semi normed group. -/ diff --git a/src/analysis/normed/group/basic.lean b/src/analysis/normed/group/basic.lean index 82370efd225e3..7b25f1a4fc002 100644 --- a/src/analysis/normed/group/basic.lean +++ b/src/analysis/normed/group/basic.lean @@ -14,6 +14,9 @@ import topology.sequences /-! # Normed (semi)groups +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define 10 classes: * `has_norm`, `has_nnnorm`: auxiliary classes endowing a type `α` with a function `norm : α → ℝ` diff --git a/src/analysis/normed/group/completion.lean b/src/analysis/normed/group/completion.lean index 8e3ba8389dc67..44087d460df0c 100644 --- a/src/analysis/normed/group/completion.lean +++ b/src/analysis/normed/group/completion.lean @@ -10,6 +10,9 @@ import topology.metric_space.completion /-! # Completion of a normed group +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove that the completion of a (semi)normed group is a normed group. ## Tags diff --git a/src/analysis/normed/group/hom.lean b/src/analysis/normed/group/hom.lean index 4954266502108..3f6b02c9df52e 100644 --- a/src/analysis/normed/group/hom.lean +++ b/src/analysis/normed/group/hom.lean @@ -8,6 +8,9 @@ import analysis.normed.group.basic /-! # Normed groups homomorphisms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file gathers definitions and elementary constructions about bounded group homomorphisms between normed (abelian) groups (abbreviated to "normed group homs"). diff --git a/src/analysis/normed/group/hom_completion.lean b/src/analysis/normed/group/hom_completion.lean index ab054fd2d501e..e3333fe0cc972 100644 --- a/src/analysis/normed/group/hom_completion.lean +++ b/src/analysis/normed/group/hom_completion.lean @@ -9,6 +9,9 @@ import analysis.normed.group.completion /-! # Completion of normed group homs +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given two (semi) normed groups `G` and `H` and a normed group hom `f : normed_add_group_hom G H`, we build and study a normed group hom `f.completion : normed_add_group_hom (completion G) (completion H)` such that the diagram diff --git a/src/analysis/normed/group/infinite_sum.lean b/src/analysis/normed/group/infinite_sum.lean index 0ab3b14942a2a..84b2afbcb9395 100644 --- a/src/analysis/normed/group/infinite_sum.lean +++ b/src/analysis/normed/group/infinite_sum.lean @@ -10,6 +10,9 @@ import topology.instances.nnreal /-! # Infinite sums in (semi)normed groups +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In a complete (semi)normed group, - `summable_iff_vanishing_norm`: a series `∑' i, f i` is summable if and only if for any `ε > 0`, diff --git a/src/analysis/normed/group/seminorm.lean b/src/analysis/normed/group/seminorm.lean index 5bc0b7f3f7854..08ea6081b6a2f 100644 --- a/src/analysis/normed/group/seminorm.lean +++ b/src/analysis/normed/group/seminorm.lean @@ -9,6 +9,9 @@ import data.real.nnreal /-! # Group seminorms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines norms and seminorms in a group. A group seminorm is a function to the reals which is positive-semidefinite and subadditive. A norm further only maps zero to zero. diff --git a/src/analysis/normed_space/indicator_function.lean b/src/analysis/normed_space/indicator_function.lean index 0214d2e6582ca..f30351b556fb7 100644 --- a/src/analysis/normed_space/indicator_function.lean +++ b/src/analysis/normed_space/indicator_function.lean @@ -9,6 +9,9 @@ import algebra.indicator_function /-! # Indicator function and norm +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains a few simple lemmas about `set.indicator` and `norm`. ## Tags diff --git a/src/analysis/subadditive.lean b/src/analysis/subadditive.lean index 6739d92a06530..c7e1d6f3311d4 100644 --- a/src/analysis/subadditive.lean +++ b/src/analysis/subadditive.lean @@ -9,6 +9,9 @@ import order.filter.archimedean /-! # Convergence of subadditive sequences +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A subadditive sequence `u : ℕ → ℝ` is a sequence satisfying `u (m + n) ≤ u m + u n` for all `m, n`. We define this notion as `subadditive u`, and prove in `subadditive.tendsto_lim` that, if `u n / n` is bounded below, then it converges to a limit (that we denote by `subadditive.lim` for diff --git a/src/category_theory/abelian/images.lean b/src/category_theory/abelian/images.lean index 5e2d2a2bf1c46..84368319191fc 100644 --- a/src/category_theory/abelian/images.lean +++ b/src/category_theory/abelian/images.lean @@ -8,6 +8,9 @@ import category_theory.limits.shapes.kernels /-! # The abelian image and coimage. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In an abelian category we usually want the image of a morphism `f` to be defined as `kernel (cokernel.π f)`, and the coimage to be defined as `cokernel (kernel.ι f)`. diff --git a/src/category_theory/abelian/non_preadditive.lean b/src/category_theory/abelian/non_preadditive.lean index 457ca077e65eb..13f11f13ec1a9 100644 --- a/src/category_theory/abelian/non_preadditive.lean +++ b/src/category_theory/abelian/non_preadditive.lean @@ -12,6 +12,9 @@ import category_theory.preadditive.basic /-! # Every non_preadditive_abelian category is preadditive +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In mathlib, we define an abelian category as a preadditive category with a zero object, kernels and cokernels, products and coproducts and in which every monomorphism and epimorphis is normal. diff --git a/src/category_theory/adjunction/comma.lean b/src/category_theory/adjunction/comma.lean index 035c60a566ff2..b4668efc61b2e 100644 --- a/src/category_theory/adjunction/comma.lean +++ b/src/category_theory/adjunction/comma.lean @@ -10,6 +10,9 @@ import category_theory.structured_arrow /-! # Properties of comma categories relating to adjunctions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file shows that for a functor `G : D ⥤ C` the data of an initial object in each `structured_arrow` category on `G` is equivalent to a left adjoint to `G`, as well as the dual. diff --git a/src/category_theory/concrete_category/reflects_isomorphisms.lean b/src/category_theory/concrete_category/reflects_isomorphisms.lean index b87261233dd51..c604560f6369f 100644 --- a/src/category_theory/concrete_category/reflects_isomorphisms.lean +++ b/src/category_theory/concrete_category/reflects_isomorphisms.lean @@ -7,6 +7,9 @@ import category_theory.concrete_category.basic import category_theory.functor.reflects_isomorphisms /-! +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A `forget₂ C D` forgetful functor between concrete categories `C` and `D` whose forgetful functors both reflect isomorphisms, itself reflects isomorphisms. -/ diff --git a/src/category_theory/is_connected.lean b/src/category_theory/is_connected.lean index 382b8f24f43e1..5c24658d6df57 100644 --- a/src/category_theory/is_connected.lean +++ b/src/category_theory/is_connected.lean @@ -11,6 +11,9 @@ import category_theory.category.ulift /-! # Connected category +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Define a connected category as a _nonempty_ category for which every functor to a discrete category is isomorphic to the constant functor. diff --git a/src/category_theory/limits/bicones.lean b/src/category_theory/limits/bicones.lean index a6a8efb89f276..ced55c90ecb55 100644 --- a/src/category_theory/limits/bicones.lean +++ b/src/category_theory/limits/bicones.lean @@ -9,6 +9,9 @@ import category_theory.fin_category /-! # Bicones +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given a category `J`, a walking `bicone J` is a category whose objects are the objects of `J` and two extra vertices `bicone.left` and `bicone.right`. The morphisms are the morphisms of `J` and `left ⟶ j`, `right ⟶ j` for each `j : J` such that `⬝ ⟶ j` and `⬝ ⟶ k` commutes with each diff --git a/src/category_theory/limits/comma.lean b/src/category_theory/limits/comma.lean index 6de3508a4374b..f687492af0777 100644 --- a/src/category_theory/limits/comma.lean +++ b/src/category_theory/limits/comma.lean @@ -12,6 +12,9 @@ import category_theory.structured_arrow /-! # Limits and colimits in comma categories +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We build limits in the comma category `comma L R` provided that the two source categories have limits and `R` preserves them. This is used to construct limits in the arrow category, structured arrow category and under diff --git a/src/category_theory/limits/cone_category.lean b/src/category_theory/limits/cone_category.lean index 21da24abd0217..34ea874977a06 100644 --- a/src/category_theory/limits/cone_category.lean +++ b/src/category_theory/limits/cone_category.lean @@ -11,6 +11,9 @@ import category_theory.limits.shapes.equivalence /-! # Limits and the category of (co)cones +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This files contains results that stem from the limit API. For the definition and the category instance of `cone`, please refer to `category_theory/limits/cones.lean`. diff --git a/src/category_theory/limits/constructions/equalizers.lean b/src/category_theory/limits/constructions/equalizers.lean index fb51d341ed810..e311be168e3fc 100644 --- a/src/category_theory/limits/constructions/equalizers.lean +++ b/src/category_theory/limits/constructions/equalizers.lean @@ -12,6 +12,9 @@ import category_theory.limits.preserves.shapes.binary_products /-! # Constructing equalizers from pullbacks and binary products. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + If a category has pullbacks and binary products, then it has equalizers. TODO: generalize universe diff --git a/src/category_theory/limits/constructions/finite_products_of_binary_products.lean b/src/category_theory/limits/constructions/finite_products_of_binary_products.lean index b85b98d01021d..ff8793b35fff6 100644 --- a/src/category_theory/limits/constructions/finite_products_of_binary_products.lean +++ b/src/category_theory/limits/constructions/finite_products_of_binary_products.lean @@ -12,6 +12,9 @@ import logic.equiv.fin /-! # Constructing finite products from binary products and terminal. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + If a category has binary products and a terminal object then it has finite products. If a functor preserves binary products and the terminal object then it preserves finite products. 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 0be2e07faa027..c022171b26cb2 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 @@ -17,6 +17,9 @@ import category_theory.limits.constructions.binary_products /-! # Constructing limits from products and equalizers. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + If a category has all products, and all equalizers, then it has all limits. Similarly, if it has all finite products, and all equalizers, then it has all finite limits. diff --git a/src/category_theory/limits/essentially_small.lean b/src/category_theory/limits/essentially_small.lean index 2c04a775c663d..99b2104c7cefd 100644 --- a/src/category_theory/limits/essentially_small.lean +++ b/src/category_theory/limits/essentially_small.lean @@ -9,6 +9,9 @@ import category_theory.essentially_small /-! # Limits over essentially small indexing categories +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + If `C` has limits of size `w` and `J` is `w`-essentially small, then `C` has limits of shape `J`. -/ diff --git a/src/category_theory/limits/kan_extension.lean b/src/category_theory/limits/kan_extension.lean index 8020306d93e2b..083ca5a6096ea 100644 --- a/src/category_theory/limits/kan_extension.lean +++ b/src/category_theory/limits/kan_extension.lean @@ -11,6 +11,9 @@ import category_theory.structured_arrow # Kan extensions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the right and left Kan extensions of a functor. They exist under the assumption that the target category has enough limits resp. colimits. diff --git a/src/category_theory/limits/lattice.lean b/src/category_theory/limits/lattice.lean index f6c4ddf059106..fce7bead9d70f 100644 --- a/src/category_theory/limits/lattice.lean +++ b/src/category_theory/limits/lattice.lean @@ -12,6 +12,9 @@ import category_theory.limits.shapes.finite_limits /-! # Limits in lattice categories are given by infimums and supremums. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ universes w u diff --git a/src/category_theory/limits/mono_coprod.lean b/src/category_theory/limits/mono_coprod.lean index 4a5b437d63efb..76dd0b27804a1 100644 --- a/src/category_theory/limits/mono_coprod.lean +++ b/src/category_theory/limits/mono_coprod.lean @@ -11,6 +11,9 @@ import category_theory.limits.shapes.zero_morphisms # Categories where inclusions into coproducts are monomorphisms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + If `C` is a category, the class `mono_coprod C` expresses that left inclusions `A ⟶ A ⨿ B` are monomorphisms when `has_coproduct A B` is satisfied. If it is so, it is shown that right inclusions are diff --git a/src/category_theory/limits/opposites.lean b/src/category_theory/limits/opposites.lean index 10307bef6915d..00f7ca6069333 100644 --- a/src/category_theory/limits/opposites.lean +++ b/src/category_theory/limits/opposites.lean @@ -11,6 +11,9 @@ import tactic.equiv_rw /-! # Limits in `C` give colimits in `Cᵒᵖ`. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We also give special cases for (co)products, (co)equalizers, and pullbacks / pushouts. diff --git a/src/category_theory/limits/preserves/filtered.lean b/src/category_theory/limits/preserves/filtered.lean index 149eba878f831..aab016a35621e 100644 --- a/src/category_theory/limits/preserves/filtered.lean +++ b/src/category_theory/limits/preserves/filtered.lean @@ -8,6 +8,9 @@ import category_theory.filtered /-! # Preservation of filtered colimits and cofiltered limits. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. Typically forgetful functors from algebraic categories preserve filtered colimits (although not general colimits). See e.g. `algebra/category/Mon/filtered_colimits`. diff --git a/src/category_theory/limits/preserves/shapes/biproducts.lean b/src/category_theory/limits/preserves/shapes/biproducts.lean index 54811cb035b01..d98231a3610dc 100644 --- a/src/category_theory/limits/preserves/shapes/biproducts.lean +++ b/src/category_theory/limits/preserves/shapes/biproducts.lean @@ -9,6 +9,9 @@ import category_theory.limits.preserves.shapes.zero /-! # Preservation of biproducts +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define the image of a (binary) bicone under a functor that preserves zero morphisms and define classes `preserves_biproduct` and `preserves_binary_biproduct`. We then diff --git a/src/category_theory/limits/preserves/shapes/zero.lean b/src/category_theory/limits/preserves/shapes/zero.lean index 999241a5ad291..643c7c51ab7b6 100644 --- a/src/category_theory/limits/preserves/shapes/zero.lean +++ b/src/category_theory/limits/preserves/shapes/zero.lean @@ -9,6 +9,9 @@ import category_theory.limits.shapes.zero_morphisms /-! # Preservation of zero objects and zero morphisms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define the class `preserves_zero_morphisms` and show basic properties. ## Main results diff --git a/src/category_theory/limits/shapes/biproducts.lean b/src/category_theory/limits/shapes/biproducts.lean index 1419e01df9a83..bef019a696072 100644 --- a/src/category_theory/limits/shapes/biproducts.lean +++ b/src/category_theory/limits/shapes/biproducts.lean @@ -10,6 +10,9 @@ import category_theory.limits.shapes.kernels /-! # Biproducts and binary biproducts +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We introduce the notion of (finite) biproducts and binary biproducts. These are slightly unusual relative to the other shapes in the library, diff --git a/src/category_theory/limits/shapes/finite_products.lean b/src/category_theory/limits/shapes/finite_products.lean index 17db52fbee3d9..701fbf5e74511 100644 --- a/src/category_theory/limits/shapes/finite_products.lean +++ b/src/category_theory/limits/shapes/finite_products.lean @@ -9,6 +9,9 @@ import category_theory.limits.shapes.products /-! # Categories with finite (co)products +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Typeclasses representing categories with (co)products over finite indexing types. -/ diff --git a/src/category_theory/limits/shapes/functor_category.lean b/src/category_theory/limits/shapes/functor_category.lean index 7b89e90b9daf2..e2d84559f7d9e 100644 --- a/src/category_theory/limits/shapes/functor_category.lean +++ b/src/category_theory/limits/shapes/functor_category.lean @@ -9,6 +9,9 @@ import category_theory.limits.functor_category /-! # If `D` has finite (co)limits, so do the functor categories `C ⥤ D`. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + These are boiler-plate instances, in their own file as neither import otherwise needs the other. -/ diff --git a/src/category_theory/limits/shapes/kernels.lean b/src/category_theory/limits/shapes/kernels.lean index 0360d36e011b0..862a36f2f68cf 100644 --- a/src/category_theory/limits/shapes/kernels.lean +++ b/src/category_theory/limits/shapes/kernels.lean @@ -8,6 +8,9 @@ import category_theory.limits.preserves.shapes.zero /-! # Kernels and cokernels +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In a category with zero morphisms, the kernel of a morphism `f : X ⟶ Y` is the equalizer of `f` and `0 : X ⟶ Y`. (Similarly the cokernel is the coequalizer.) diff --git a/src/category_theory/limits/shapes/normal_mono/basic.lean b/src/category_theory/limits/shapes/normal_mono/basic.lean index c64ff3a9daa74..c0c06beb4659d 100644 --- a/src/category_theory/limits/shapes/normal_mono/basic.lean +++ b/src/category_theory/limits/shapes/normal_mono/basic.lean @@ -10,6 +10,9 @@ import category_theory.limits.preserves.basic /-! # Definitions and basic properties of normal monomorphisms and epimorphisms. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A normal monomorphism is a morphism that is the kernel of some other morphism. We give the construction `normal_mono → regular_mono` (`category_theory.normal_mono.regular_mono`) diff --git a/src/category_theory/limits/shapes/normal_mono/equalizers.lean b/src/category_theory/limits/shapes/normal_mono/equalizers.lean index 2fd7929120b71..1c620830185d1 100644 --- a/src/category_theory/limits/shapes/normal_mono/equalizers.lean +++ b/src/category_theory/limits/shapes/normal_mono/equalizers.lean @@ -9,6 +9,9 @@ import category_theory.limits.shapes.finite_products /-! # Normal mono categories with finite products and kernels have all equalizers. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This, and the dual result, are used in the development of abelian categories. -/ diff --git a/src/category_theory/limits/shapes/zero_morphisms.lean b/src/category_theory/limits/shapes/zero_morphisms.lean index c99d2118ad41f..a4c8d04ca24b4 100644 --- a/src/category_theory/limits/shapes/zero_morphisms.lean +++ b/src/category_theory/limits/shapes/zero_morphisms.lean @@ -12,6 +12,9 @@ import category_theory.limits.shapes.zero_objects /-! # Zero morphisms and zero objects +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A category "has zero morphisms" if there is a designated "zero morphism" in each morphism space, and compositions of zero morphisms with anything give the zero morphism. (Notice this is extra structure, not merely a property.) diff --git a/src/category_theory/limits/small_complete.lean b/src/category_theory/limits/small_complete.lean index 5eed263d4d608..c61016db8abe7 100644 --- a/src/category_theory/limits/small_complete.lean +++ b/src/category_theory/limits/small_complete.lean @@ -10,6 +10,9 @@ import set_theory.cardinal.basic /-! # Any small complete category is a preorder +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We show that any small category which has all (small) limits is a preorder: In particular, we show that if a small category `C` in universe `u` has products of size `u`, then for any `X Y : C` there is at most one morphism `X ⟶ Y`. diff --git a/src/category_theory/limits/types.lean b/src/category_theory/limits/types.lean index 365ae72bf946c..e9e7959d7b6b9 100644 --- a/src/category_theory/limits/types.lean +++ b/src/category_theory/limits/types.lean @@ -10,6 +10,9 @@ import tactic.equiv_rw /-! # Limits in the category of types. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We show that the category of types has all (co)limits, by providing the usual concrete models. We also give a characterisation of filtered colimits in `Type`, via diff --git a/src/category_theory/linear/basic.lean b/src/category_theory/linear/basic.lean index cd01f81bef101..2f1106b55a277 100644 --- a/src/category_theory/linear/basic.lean +++ b/src/category_theory/linear/basic.lean @@ -11,6 +11,9 @@ import algebra.algebra.basic /-! # Linear categories +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + 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. diff --git a/src/category_theory/linear/functor_category.lean b/src/category_theory/linear/functor_category.lean index cace39f145bce..594349a2a5381 100644 --- a/src/category_theory/linear/functor_category.lean +++ b/src/category_theory/linear/functor_category.lean @@ -9,6 +9,9 @@ import category_theory.linear.basic /-! # Linear structure on functor categories +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + If `C` and `D` are categories and `D` is `R`-linear, then `C ⥤ D` is also `R`-linear. diff --git a/src/category_theory/monoidal/End.lean b/src/category_theory/monoidal/End.lean index f12b646e85e4c..502a540e79c6e 100644 --- a/src/category_theory/monoidal/End.lean +++ b/src/category_theory/monoidal/End.lean @@ -8,6 +8,9 @@ import category_theory.monoidal.functor /-! # Endofunctors as a monoidal category. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We give the monoidal category structure on `C ⥤ C`, and show that when `C` itself is monoidal, it embeds via a monoidal functor into `C ⥤ C`. diff --git a/src/category_theory/preadditive/additive_functor.lean b/src/category_theory/preadditive/additive_functor.lean index 7b8080af28d67..34347aa14e68c 100644 --- a/src/category_theory/preadditive/additive_functor.lean +++ b/src/category_theory/preadditive/additive_functor.lean @@ -11,6 +11,9 @@ import category_theory.preadditive.functor_category /-! # Additive Functors +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A functor between two preadditive categories is called *additive* provided that the induced map on hom types is a morphism of abelian groups. diff --git a/src/category_theory/preadditive/basic.lean b/src/category_theory/preadditive/basic.lean index e2925475cc225..8bf3044d5c0c5 100644 --- a/src/category_theory/preadditive/basic.lean +++ b/src/category_theory/preadditive/basic.lean @@ -12,6 +12,9 @@ import category_theory.limits.shapes.kernels /-! # Preadditive categories +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A preadditive category is a category in which `X ⟶ Y` is an abelian group in such a way that composition of morphisms is linear in both variables. diff --git a/src/category_theory/preadditive/biproducts.lean b/src/category_theory/preadditive/biproducts.lean index 90d9aa815d9f7..b6f15e50a9d78 100644 --- a/src/category_theory/preadditive/biproducts.lean +++ b/src/category_theory/preadditive/biproducts.lean @@ -14,6 +14,9 @@ import tactic.abel /-! # Basic facts about biproducts in preadditive categories. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In (or between) preadditive categories, * Any biproduct satisfies the equality diff --git a/src/category_theory/preadditive/functor_category.lean b/src/category_theory/preadditive/functor_category.lean index 602be9aebf3e6..6117d69648326 100644 --- a/src/category_theory/preadditive/functor_category.lean +++ b/src/category_theory/preadditive/functor_category.lean @@ -9,6 +9,9 @@ import category_theory.preadditive.basic /-! # Preadditive structure on functor categories +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + If `C` and `D` are categories and `D` is preadditive, then `C ⥤ D` is also preadditive. diff --git a/src/combinatorics/hindman.lean b/src/combinatorics/hindman.lean index ca9e61f6c7bc3..2a7da25298b50 100644 --- a/src/combinatorics/hindman.lean +++ b/src/combinatorics/hindman.lean @@ -10,6 +10,9 @@ import data.stream.init /-! # Hindman's theorem on finite sums +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We prove Hindman's theorem on finite sums, using idempotent ultrafilters. Given an infinite sequence `a₀, a₁, a₂, …` of positive integers, the set `FS(a₀, …)` is the set diff --git a/src/control/bifunctor.lean b/src/control/bifunctor.lean index 356eab1bda53a..fae259c06f95a 100644 --- a/src/control/bifunctor.lean +++ b/src/control/bifunctor.lean @@ -9,6 +9,9 @@ import data.sum.basic /-! # Functors with two arguments +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines bifunctors. A bifunctor is a function `F : Type* → Type* → Type*` along with a bimap which turns `F α β` into diff --git a/src/control/bitraversable/basic.lean b/src/control/bitraversable/basic.lean index c91046710290f..2d9b48ea6a9f7 100644 --- a/src/control/bitraversable/basic.lean +++ b/src/control/bitraversable/basic.lean @@ -9,6 +9,9 @@ import control.traversable.basic /-! # Bitraversable type class +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Type class for traversing bifunctors. Simple examples of `bitraversable` are `prod` and `sum`. A more elaborate example is diff --git a/src/data/complex/basic.lean b/src/data/complex/basic.lean index dfbe4449da6fb..887274d0b7c8e 100644 --- a/src/data/complex/basic.lean +++ b/src/data/complex/basic.lean @@ -8,6 +8,9 @@ import data.real.sqrt /-! # The complex numbers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The complex numbers are modelled as ℝ^2 in the obvious way and it is shown that they form a field of characteristic zero. The result that the complex numbers are algebraically closed, see `field_theory.algebraic_closure`. diff --git a/src/data/nat/choose/cast.lean b/src/data/nat/choose/cast.lean index 907dc455792bb..d89a3b10422c1 100644 --- a/src/data/nat/choose/cast.lean +++ b/src/data/nat/choose/cast.lean @@ -9,6 +9,9 @@ import data.nat.factorial.cast /-! # Cast of binomial coefficients +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file allows calculating the binomial coefficient `a.choose b` as an element of a division ring of characteristic `0`. -/ diff --git a/src/data/nat/choose/vandermonde.lean b/src/data/nat/choose/vandermonde.lean index b9c0db02e3147..909280a1c668b 100644 --- a/src/data/nat/choose/vandermonde.lean +++ b/src/data/nat/choose/vandermonde.lean @@ -11,6 +11,9 @@ import data.nat.choose.basic # Vandermonde's identity +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove Vandermonde's identity (`nat.add_choose_eq`): `(m + n).choose k = ∑ (ij : ℕ × ℕ) in antidiagonal k, m.choose ij.1 * n.choose ij.2` diff --git a/src/data/nat/factorial/cast.lean b/src/data/nat/factorial/cast.lean index 9acc2393ce578..73fbff0a63894 100644 --- a/src/data/nat/factorial/cast.lean +++ b/src/data/nat/factorial/cast.lean @@ -8,6 +8,9 @@ import ring_theory.polynomial.pochhammer /-! # Cast of factorials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file allows calculating factorials (including ascending and descending ones) as elements of a semiring. diff --git a/src/data/polynomial/basic.lean b/src/data/polynomial/basic.lean index 7b8027fd78cf6..7334c8cdcc10c 100644 --- a/src/data/polynomial/basic.lean +++ b/src/data/polynomial/basic.lean @@ -9,6 +9,9 @@ import data.finset.sort /-! # Theory of univariate polynomials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines `polynomial R`, the type of univariate polynomials over the semiring `R`, builds a semiring structure on it, and gives basic definitions that are expanded in other files in this directory. diff --git a/src/data/polynomial/cardinal.lean b/src/data/polynomial/cardinal.lean index 6bd5486e7f972..a398efb2fb04b 100644 --- a/src/data/polynomial/cardinal.lean +++ b/src/data/polynomial/cardinal.lean @@ -8,6 +8,9 @@ import set_theory.cardinal.ordinal /-! # Cardinality of Polynomial Ring +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The reuslt in this file is that the cardinality of `R[X]` is at most the maximum of `#R` and `ℵ₀`. -/ diff --git a/src/data/polynomial/coeff.lean b/src/data/polynomial/coeff.lean index d08b4e42cd584..8bcb8e2638580 100644 --- a/src/data/polynomial/coeff.lean +++ b/src/data/polynomial/coeff.lean @@ -11,6 +11,9 @@ import data.nat.choose.sum /-! # Theory of univariate polynomials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The theorems include formulas for computing coefficients, such as `coeff_add`, `coeff_sum`, `coeff_mul` diff --git a/src/data/polynomial/degree/definitions.lean b/src/data/polynomial/degree/definitions.lean index 3aee2dc61ef71..895f51d0483e2 100644 --- a/src/data/polynomial/degree/definitions.lean +++ b/src/data/polynomial/degree/definitions.lean @@ -11,6 +11,9 @@ import data.polynomial.coeff /-! # Theory of univariate polynomials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The definitions include `degree`, `monic`, `leading_coeff` diff --git a/src/data/polynomial/degree/lemmas.lean b/src/data/polynomial/degree/lemmas.lean index f9f10db9a9563..43ca73a90d635 100644 --- a/src/data/polynomial/degree/lemmas.lean +++ b/src/data/polynomial/degree/lemmas.lean @@ -8,6 +8,9 @@ import data.polynomial.eval /-! # Theory of degrees of polynomials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Some of the main results include - `nat_degree_comp_le` : The degree of the composition is at most the product of degrees diff --git a/src/data/polynomial/degree/trailing_degree.lean b/src/data/polynomial/degree/trailing_degree.lean index 3903abea51ee5..b78c13868af37 100644 --- a/src/data/polynomial/degree/trailing_degree.lean +++ b/src/data/polynomial/degree/trailing_degree.lean @@ -9,6 +9,9 @@ import data.polynomial.degree.definitions /-! # Trailing degree of univariate polynomials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions * `trailing_degree p`: the multiplicity of `X` in the polynomial `p` diff --git a/src/data/polynomial/derivative.lean b/src/data/polynomial/derivative.lean index 170c4700a4fba..7f8bccce59b7f 100644 --- a/src/data/polynomial/derivative.lean +++ b/src/data/polynomial/derivative.lean @@ -9,6 +9,9 @@ import data.polynomial.eval /-! # The derivative map on polynomials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions * `polynomial.derivative`: The formal derivative of polynomials, expressed as a linear map. diff --git a/src/data/polynomial/erase_lead.lean b/src/data/polynomial/erase_lead.lean index df4431851b68b..0b6b93a170f37 100644 --- a/src/data/polynomial/erase_lead.lean +++ b/src/data/polynomial/erase_lead.lean @@ -9,6 +9,9 @@ import data.polynomial.degree.definitions /-! # Erase the leading term of a univariate polynomial +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Definition * `erase_lead f`: the polynomial `f - leading term of f` diff --git a/src/data/polynomial/eval.lean b/src/data/polynomial/eval.lean index 3a39fd1b3db4e..b57008b992ac0 100644 --- a/src/data/polynomial/eval.lean +++ b/src/data/polynomial/eval.lean @@ -9,6 +9,9 @@ import data.polynomial.induction /-! # Theory of univariate polynomials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The main defs here are `eval₂`, `eval`, and `map`. We give several lemmas about their interaction with each other and with module operations. -/ diff --git a/src/data/polynomial/induction.lean b/src/data/polynomial/induction.lean index ff1629c69def6..1e3bc53228020 100644 --- a/src/data/polynomial/induction.lean +++ b/src/data/polynomial/induction.lean @@ -9,6 +9,9 @@ import data.polynomial.basic /-! # Induction on polynomials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains lemmas dealing with different flavours of induction on polynomials. See also `data/polynomial/inductions.lean` (with an `s`!). diff --git a/src/data/polynomial/inductions.lean b/src/data/polynomial/inductions.lean index dea3738a0f773..7dd0f627543db 100644 --- a/src/data/polynomial/inductions.lean +++ b/src/data/polynomial/inductions.lean @@ -10,6 +10,9 @@ import data.polynomial.induction /-! # Induction on polynomials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains lemmas dealing with different flavours of induction on polynomials. -/ diff --git a/src/data/polynomial/monic.lean b/src/data/polynomial/monic.lean index b2cf182f7f54d..e04a1006eb560 100644 --- a/src/data/polynomial/monic.lean +++ b/src/data/polynomial/monic.lean @@ -9,6 +9,9 @@ import algebra.regular.smul /-! # Theory of monic polynomials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We give several tools for proving that polynomials are monic, e.g. `monic.mul`, `monic.map`, `monic.pow`. -/ diff --git a/src/data/polynomial/monomial.lean b/src/data/polynomial/monomial.lean index eb8f32669e6bb..27bbac66f924d 100644 --- a/src/data/polynomial/monomial.lean +++ b/src/data/polynomial/monomial.lean @@ -8,6 +8,9 @@ import data.polynomial.basic /-! # Univariate monomials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Preparatory lemmas for degree_basic. -/ diff --git a/src/data/polynomial/reverse.lean b/src/data/polynomial/reverse.lean index eb328346d277f..037c15d2c939e 100644 --- a/src/data/polynomial/reverse.lean +++ b/src/data/polynomial/reverse.lean @@ -10,6 +10,9 @@ import data.polynomial.eval /-! # Reverse of a univariate polynomial +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The main definition is `reverse`. Applying `reverse` to a polynomial `f : R[X]` produces the polynomial with a reversed list of coefficients, equivalent to `X^f.nat_degree * f(1/X)`. diff --git a/src/data/real/sqrt.lean b/src/data/real/sqrt.lean index 27e44303b38c8..c8f1337bd8715 100644 --- a/src/data/real/sqrt.lean +++ b/src/data/real/sqrt.lean @@ -10,6 +10,9 @@ import tactic.positivity /-! # Square root of a real number +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define * `nnreal.sqrt` to be the square root of a nonnegative real number. diff --git a/src/data/tree.lean b/src/data/tree.lean index 50c55cd712a64..d2e2981efb571 100644 --- a/src/data/tree.lean +++ b/src/data/tree.lean @@ -10,6 +10,9 @@ import order.basic /-! # Binary tree +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Provides binary tree storage for values of any type, with O(lg n) retrieval. See also `data.rbtree` for red-black trees - this version allows more operations to be defined and is better suited for in-kernel computation. diff --git a/src/linear_algebra/affine_space/slope.lean b/src/linear_algebra/affine_space/slope.lean index a279a91d9bf50..768a9857022b2 100644 --- a/src/linear_algebra/affine_space/slope.lean +++ b/src/linear_algebra/affine_space/slope.lean @@ -9,6 +9,9 @@ import tactic.field_simp /-! # Slope of a function +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define the slope of a function `f : k → PE` taking values in an affine space over `k` and prove some basic theorems about `slope`. The `slope` function naturally appears in the Mean Value Theorem, and in the proof of the fact that a function with nonnegative second derivative on an diff --git a/src/logic/equiv/functor.lean b/src/logic/equiv/functor.lean index 35b625f2b8f35..54d3dfeed9eb3 100644 --- a/src/logic/equiv/functor.lean +++ b/src/logic/equiv/functor.lean @@ -9,6 +9,9 @@ import logic.equiv.defs /-! # Functor and bifunctors can be applied to `equiv`s. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define ```lean def functor.map_equiv (f : Type u → Type v) [functor f] [is_lawful_functor f] : diff --git a/src/number_theory/basic.lean b/src/number_theory/basic.lean index 0accec3918f5b..17b26e3f795c8 100644 --- a/src/number_theory/basic.lean +++ b/src/number_theory/basic.lean @@ -10,6 +10,9 @@ import ring_theory.ideal.quotient /-! # Basic results in number theory +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file should contain basic results in number theory. So far, it only contains the essential lemma in the construction of the ring of Witt vectors. diff --git a/src/order/jordan_holder.lean b/src/order/jordan_holder.lean index 477d33bd61430..81b4c25040e92 100644 --- a/src/order/jordan_holder.lean +++ b/src/order/jordan_holder.lean @@ -11,6 +11,9 @@ import data.fintype.card /-! # Jordan-Hölder Theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves the Jordan Hölder theorem for a `jordan_holder_lattice`, a class also defined in this file. Examples of `jordan_holder_lattice` include `subgroup G` if `G` is a group, and `submodule R M` if `M` is an `R`-module. Using this approach the theorem need not be proved diff --git a/src/ring_theory/adjoin/basic.lean b/src/ring_theory/adjoin/basic.lean index 042265be90932..1e8853cbe9fbf 100644 --- a/src/ring_theory/adjoin/basic.lean +++ b/src/ring_theory/adjoin/basic.lean @@ -11,6 +11,9 @@ import linear_algebra.finsupp /-! # Adjoining elements to form subalgebras +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file develops the basic theory of subalgebras of an R-algebra generated by a set of elements. A basic interface for `adjoin` is set up. diff --git a/src/ring_theory/ideal/operations.lean b/src/ring_theory/ideal/operations.lean index 1288fbfd43702..aef44435a9b5a 100644 --- a/src/ring_theory/ideal/operations.lean +++ b/src/ring_theory/ideal/operations.lean @@ -12,6 +12,9 @@ import ring_theory.ideal.basic import ring_theory.non_zero_divisors /-! # More operations on modules and ideals + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ universes u v w x diff --git a/src/ring_theory/ideal/quotient.lean b/src/ring_theory/ideal/quotient.lean index 2ebb7ee7072e5..cda8ece37d175 100644 --- a/src/ring_theory/ideal/quotient.lean +++ b/src/ring_theory/ideal/quotient.lean @@ -12,6 +12,9 @@ import tactic.fin_cases /-! # Ideal quotients +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines ideal quotients as a special case of submodule quotients and proves some basic results about these quotients. diff --git a/src/ring_theory/localization/basic.lean b/src/ring_theory/localization/basic.lean index f26c26a139e19..06c272a17bd50 100644 --- a/src/ring_theory/localization/basic.lean +++ b/src/ring_theory/localization/basic.lean @@ -13,6 +13,9 @@ import tactic.ring_exp /-! # Localizations of commutative rings +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We characterize the localization of a commutative ring `R` at a submonoid `M` up to isomorphism; that is, a commutative ring `S` is the localization of `R` at `M` iff we can find a ring homomorphism `f : R →+* S` satisfying 3 properties: diff --git a/src/ring_theory/nilpotent.lean b/src/ring_theory/nilpotent.lean index 775b5888b7d28..1795de367c5f9 100644 --- a/src/ring_theory/nilpotent.lean +++ b/src/ring_theory/nilpotent.lean @@ -10,6 +10,9 @@ import ring_theory.ideal.operations /-! # Nilpotent elements +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions * `is_nilpotent` diff --git a/src/ring_theory/polynomial/pochhammer.lean b/src/ring_theory/polynomial/pochhammer.lean index 6907ad181a46d..2a9d59baac6c1 100644 --- a/src/ring_theory/polynomial/pochhammer.lean +++ b/src/ring_theory/polynomial/pochhammer.lean @@ -9,6 +9,9 @@ import data.polynomial.eval /-! # The Pochhammer polynomials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define and prove some basic relations about `pochhammer S n : S[X] := X * (X + 1) * ... * (X + n - 1)` which is also known as the rising factorial. A version of this definition diff --git a/src/topology/algebra/affine.lean b/src/topology/algebra/affine.lean index 4f1934a0067c8..6c36789010ae4 100644 --- a/src/topology/algebra/affine.lean +++ b/src/topology/algebra/affine.lean @@ -10,6 +10,9 @@ import topology.algebra.mul_action /-! # Topological properties of affine spaces and maps +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + For now, this contains only a few facts regarding the continuity of affine maps in the special case when the point space and vector space are the same. diff --git a/src/topology/algebra/localization.lean b/src/topology/algebra/localization.lean index 068d93e424869..28cf23992b53b 100644 --- a/src/topology/algebra/localization.lean +++ b/src/topology/algebra/localization.lean @@ -10,6 +10,9 @@ import topology.algebra.ring.basic # Localization of topological rings +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The topological localization of a topological commutative ring `R` at a submonoid `M` is the ring `localization M` endowed with the final ring topology of the natural homomorphism sending `x : R` to the equivalence class of `(x, 1)` in the localization of `R` at a `M`. diff --git a/src/topology/algebra/ring/ideal.lean b/src/topology/algebra/ring/ideal.lean index 90d873889ef38..a39462b874272 100644 --- a/src/topology/algebra/ring/ideal.lean +++ b/src/topology/algebra/ring/ideal.lean @@ -8,6 +8,9 @@ import ring_theory.ideal.quotient /-! # Ideals and quotients of topological rings +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define `ideal.closure` to be the topological closure of an ideal in a topological ring. We also define a `topological_space` structure on the quotient of a topological ring by an ideal and prove that the quotient is a topological ring. diff --git a/src/topology/fiber_bundle/trivialization.lean b/src/topology/fiber_bundle/trivialization.lean index b69408370640f..84a15ebb82f14 100644 --- a/src/topology/fiber_bundle/trivialization.lean +++ b/src/topology/fiber_bundle/trivialization.lean @@ -10,6 +10,9 @@ import topology.local_homeomorph /-! # Trivializations +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions ### Basic definitions diff --git a/src/topology/instances/ennreal.lean b/src/topology/instances/ennreal.lean index ecdd50d451718..4937ad6818752 100644 --- a/src/topology/instances/ennreal.lean +++ b/src/topology/instances/ennreal.lean @@ -11,6 +11,9 @@ import topology.metric_space.lipschitz /-! # Extended non-negative reals + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ noncomputable theory diff --git a/src/topology/metric_space/completion.lean b/src/topology/metric_space/completion.lean index 566a0eff7aa80..1045b5d7de514 100644 --- a/src/topology/metric_space/completion.lean +++ b/src/topology/metric_space/completion.lean @@ -10,6 +10,9 @@ import topology.instances.real /-! # The completion of a metric space +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Completion of uniform spaces are already defined in `topology.uniform_space.completion`. We show here that the uniform space completion of a metric space inherits a metric space structure, by extending the distance to the completion and checking that it is indeed a distance, and that diff --git a/src/topology/metric_space/gluing.lean b/src/topology/metric_space/gluing.lean index 9071591df748d..8f35142bced4f 100644 --- a/src/topology/metric_space/gluing.lean +++ b/src/topology/metric_space/gluing.lean @@ -8,6 +8,9 @@ import topology.metric_space.isometry /-! # Metric space gluing +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Gluing two metric spaces along a common subset. Formally, we are given ``` diff --git a/src/topology/metric_space/isometric_smul.lean b/src/topology/metric_space/isometric_smul.lean index f734a3402fafa..01e41594ac879 100644 --- a/src/topology/metric_space/isometric_smul.lean +++ b/src/topology/metric_space/isometric_smul.lean @@ -8,6 +8,9 @@ import topology.metric_space.isometry /-! # Group actions by isometries +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define two typeclasses: - `has_isometric_smul M X` says that `M` multiplicatively acts on a (pseudo extended) metric space diff --git a/src/topology/metric_space/isometry.lean b/src/topology/metric_space/isometry.lean index 1a82febfe4a0b..ef29d1dcb3200 100644 --- a/src/topology/metric_space/isometry.lean +++ b/src/topology/metric_space/isometry.lean @@ -9,6 +9,9 @@ import topology.metric_space.antilipschitz /-! # Isometries +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define isometries, i.e., maps between emetric spaces that preserve the edistance (on metric spaces, these are exactly the maps that preserve distances), and prove their basic properties. We also introduce isometric bijections. diff --git a/src/topology/uniform_space/compare_reals.lean b/src/topology/uniform_space/compare_reals.lean index c8fd75c649d96..ded30248ad255 100644 --- a/src/topology/uniform_space/compare_reals.lean +++ b/src/topology/uniform_space/compare_reals.lean @@ -11,6 +11,9 @@ import topology.uniform_space.completion /-! # Comparison of Cauchy reals and Bourbaki reals +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In `data.real.basic` real numbers are defined using the so called Cauchy construction (although it is due to Georg Cantor). More precisely, this construction applies to commutative rings equipped with an absolute value with values in a linear ordered field. diff --git a/src/topology/unit_interval.lean b/src/topology/unit_interval.lean index 49f5f358b8253..0990b7e7a59dd 100644 --- a/src/topology/unit_interval.lean +++ b/src/topology/unit_interval.lean @@ -11,6 +11,9 @@ import data.set.intervals.instances /-! # The unit interval, as a topological space +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Use `open_locale unit_interval` to turn on the notation `I := set.Icc (0 : ℝ) (1 : ℝ)`. We provide basic instances, as well as a custom tactic for discharging From 2196ab363eb097c008d4497125e0dde23fb36db2 Mon Sep 17 00:00:00 2001 From: Michael Stoll Date: Mon, 13 Mar 2023 12:12:30 +0000 Subject: [PATCH 0620/1291] refactor(algebra/group/basic): rework lemmas on inv and neg (#17483) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR adds the following lemma (and its additive equivalent). ```lean theorem inv_eq_iff_eq_inv : a⁻¹ = b ↔ a = b⁻¹ ``` and removes `eq_inv_of_eq_inv`, `eq_inv_iff_eq_inv` and `inv_eq_iff_inv_eq` (and their additive equivalents). --- src/algebra/group/basic.lean | 14 +++----------- src/algebra/group_with_zero/basic.lean | 2 +- src/algebra/order/group/abs.lean | 4 ++-- src/algebra/order/to_interval_mod.lean | 2 +- src/algebra/periodic.lean | 4 ++-- src/algebraic_geometry/elliptic_curve/point.lean | 2 +- src/analysis/normed/group/quotient.lean | 2 +- src/analysis/special_functions/gamma.lean | 4 ++-- src/analysis/special_functions/log/basic.lean | 4 ++-- .../special_functions/trigonometric/angle.lean | 4 ++-- src/data/int/basic.lean | 7 +++---- src/data/num/lemmas.lean | 6 +++--- src/data/polynomial/mirror.lean | 4 ++-- src/data/real/conjugate_exponents.lean | 4 ++-- src/data/real/ereal.lean | 6 +++--- src/data/real/golden_ratio.lean | 2 +- src/data/set/pointwise/interval.lean | 2 +- src/field_theory/finite/basic.lean | 2 +- src/geometry/euclidean/angle/oriented/basic.lean | 4 ++-- .../euclidean/angle/oriented/rotation.lean | 2 +- src/linear_algebra/affine_space/midpoint.lean | 2 +- .../clifford_algebra/even_equiv.lean | 4 ++-- src/linear_algebra/dimension.lean | 2 +- .../integral/divergence_theorem.lean | 2 +- src/number_theory/bernoulli.lean | 4 ++-- .../legendre_symbol/quadratic_reciprocity.lean | 2 +- src/number_theory/modular.lean | 2 +- src/probability/martingale/borel_cantelli.lean | 2 +- src/ring_theory/algebraic.lean | 2 +- src/ring_theory/int/basic.lean | 2 +- src/ring_theory/valuation/basic.lean | 8 ++++---- src/ring_theory/valuation/valuation_subring.lean | 2 +- src/ring_theory/witt_vector/frobenius.lean | 2 +- src/set_theory/game/pgame.lean | 2 +- 34 files changed, 55 insertions(+), 64 deletions(-) diff --git a/src/algebra/group/basic.lean b/src/algebra/group/basic.lean index 9d06c6de2b904..04c71bc61da11 100644 --- a/src/algebra/group/basic.lean +++ b/src/algebra/group/basic.lean @@ -191,16 +191,8 @@ inv_involutive.injective @[simp, to_additive] theorem inv_inj {a b : G} : a⁻¹ = b⁻¹ ↔ a = b := inv_injective.eq_iff @[to_additive] -lemma eq_inv_of_eq_inv (h : a = b⁻¹) : b = a⁻¹ := -by simp [h] - -@[to_additive] -theorem eq_inv_iff_eq_inv : a = b⁻¹ ↔ b = a⁻¹ := -⟨eq_inv_of_eq_inv, eq_inv_of_eq_inv⟩ - -@[to_additive] -theorem inv_eq_iff_inv_eq : a⁻¹ = b ↔ b⁻¹ = a := -eq_comm.trans $ eq_inv_iff_eq_inv.trans eq_comm +theorem inv_eq_iff_eq_inv : a⁻¹ = b ↔ a = b⁻¹ := +⟨λ h, h ▸ (inv_inv a).symm, λ h, h.symm ▸ inv_inv b⟩ variables (G) @@ -399,7 +391,7 @@ theorem mul_eq_one_iff_eq_inv : a * b = 1 ↔ a = b⁻¹ := @[to_additive] theorem mul_eq_one_iff_inv_eq : a * b = 1 ↔ a⁻¹ = b := -by rw [mul_eq_one_iff_eq_inv, eq_inv_iff_eq_inv, eq_comm] +by rw [mul_eq_one_iff_eq_inv, inv_eq_iff_eq_inv] @[to_additive] theorem eq_inv_iff_mul_eq_one : a = b⁻¹ ↔ a * b = 1 := diff --git a/src/algebra/group_with_zero/basic.lean b/src/algebra/group_with_zero/basic.lean index a397f82c1f95a..457883ae90ddd 100644 --- a/src/algebra/group_with_zero/basic.lean +++ b/src/algebra/group_with_zero/basic.lean @@ -286,7 +286,7 @@ lemma one_div_ne_zero {a : G₀} (h : a ≠ 0) : 1 / a ≠ 0 := by simpa only [one_div] using inv_ne_zero h @[simp] lemma inv_eq_zero {a : G₀} : a⁻¹ = 0 ↔ a = 0 := -by rw [inv_eq_iff_inv_eq, inv_zero, eq_comm] +by rw [inv_eq_iff_eq_inv, inv_zero] @[simp] lemma zero_eq_inv {a : G₀} : 0 = a⁻¹ ↔ 0 = a := eq_comm.trans $ inv_eq_zero.trans eq_comm diff --git a/src/algebra/order/group/abs.lean b/src/algebra/order/group/abs.lean index da98dd2c757a7..890733166db8e 100644 --- a/src/algebra/order/group/abs.lean +++ b/src/algebra/order/group/abs.lean @@ -62,13 +62,13 @@ begin end lemma eq_or_eq_neg_of_abs_eq {a b : α} (h : |a| = b) : a = b ∨ a = -b := -by simpa only [← h, eq_comm, eq_neg_iff_eq_neg] using abs_choice a +by simpa only [← h, eq_comm, neg_eq_iff_eq_neg] using abs_choice a lemma abs_eq_abs {a b : α} : |a| = |b| ↔ a = b ∨ a = -b := begin refine ⟨λ h, _, λ h, _⟩, { obtain rfl | rfl := eq_or_eq_neg_of_abs_eq h; - simpa only [neg_eq_iff_neg_eq, neg_inj, or.comm, @eq_comm _ (-b)] using abs_choice b }, + simpa only [neg_eq_iff_eq_neg, neg_inj, or.comm] using abs_choice b }, { cases h; simp only [h, abs_neg] }, end diff --git a/src/algebra/order/to_interval_mod.lean b/src/algebra/order/to_interval_mod.lean index 3b484635db0c3..b153c053b0aa4 100644 --- a/src/algebra/order/to_interval_mod.lean +++ b/src/algebra/order/to_interval_mod.lean @@ -320,7 +320,7 @@ lemma to_Ico_div_neg (a : α) {b : α} (hb : 0 < b) (x : α) : begin suffices : to_Ico_div a hb (-x) = -(to_Ioc_div (-(a + b)) hb x), { rwa [neg_add, ←sub_eq_add_neg, ←to_Ioc_div_add_right', to_Ioc_div_add_right] at this }, - rw [eq_neg_iff_eq_neg, eq_comm], + rw [← neg_eq_iff_eq_neg], apply eq_to_Ioc_div_of_sub_zsmul_mem_Ioc, obtain ⟨hc, ho⟩ := sub_to_Ico_div_zsmul_mem_Ico a hb (-x), rw [←neg_lt_neg_iff, neg_sub' (-x), neg_neg, ←neg_smul] at ho, diff --git a/src/algebra/periodic.lean b/src/algebra/periodic.lean index fba880b221155..d5bd16c586aae 100644 --- a/src/algebra/periodic.lean +++ b/src/algebra/periodic.lean @@ -338,7 +338,7 @@ funext h protected lemma antiperiodic.funext' [has_add α] [has_involutive_neg β] (h : antiperiodic f c) : (λ x, -f (x + c)) = f := -(eq_neg_iff_eq_neg.mp h.funext).symm +neg_eq_iff_eq_neg.mpr h.funext /-- If a function is `antiperiodic` with antiperiod `c`, then it is also `periodic` with period `2 * c`. -/ @@ -373,7 +373,7 @@ lemma antiperiodic.int_odd_mul_antiperiodic [ring α] [has_involutive_neg β] lemma antiperiodic.sub_eq [add_group α] [has_involutive_neg β] (h : antiperiodic f c) (x : α) : f (x - c) = -f x := -by simp only [eq_neg_iff_eq_neg.mp (h (x - c)), sub_add_cancel] +by rw [← neg_eq_iff_eq_neg, ← h (x - c), sub_add_cancel] lemma antiperiodic.sub_eq' [add_comm_group α] [has_neg β] (h : antiperiodic f c) : f (c - x) = -f (-x) := diff --git a/src/algebraic_geometry/elliptic_curve/point.lean b/src/algebraic_geometry/elliptic_curve/point.lean index 665ff960c8798..eee73fcbb56f5 100644 --- a/src/algebraic_geometry/elliptic_curve/point.lean +++ b/src/algebraic_geometry/elliptic_curve/point.lean @@ -512,7 +512,7 @@ namespace point begin rcases ⟨P, Q⟩ with ⟨_ | @⟨x₁, y₁, _⟩, _ | @⟨x₂, y₂, _⟩⟩, any_goals { refl }, - { rw [zero_def, zero_add, eq_neg_iff_eq_neg, neg_zero] }, + { rw [zero_def, zero_add, ← neg_eq_iff_eq_neg, neg_zero, eq_comm], }, { simp only [neg_some], split, { intro h, diff --git a/src/analysis/normed/group/quotient.lean b/src/analysis/normed/group/quotient.lean index e6e27628f7055..1c71d8f281869 100644 --- a/src/analysis/normed/group/quotient.lean +++ b/src/analysis/normed/group/quotient.lean @@ -134,7 +134,7 @@ begin rw ← norm_neg, exact ⟨-m, by simp only [(mk' S).map_neg, set.mem_set_of_eq], rfl⟩ }, { rintros ⟨m, hm : mk' S m = -x, rfl⟩, - exact ⟨-m, by simpa [eq_comm] using eq_neg_iff_eq_neg.mp ((mk'_apply _ _).symm.trans hm)⟩ } + exact ⟨-m, by simpa using neg_eq_iff_eq_neg.mpr ((mk'_apply _ _).symm.trans hm)⟩ } end lemma quotient_norm_sub_rev {S : add_subgroup M} (x y : M ⧸ S) : ‖x - y‖ = ‖y - x‖ := diff --git a/src/analysis/special_functions/gamma.lean b/src/analysis/special_functions/gamma.lean index 358b1b428b146..964fc1cc8908e 100644 --- a/src/analysis/special_functions/gamma.lean +++ b/src/analysis/special_functions/gamma.lean @@ -1453,8 +1453,8 @@ begin rw [←neg_eq_zero, ←complex.sin_neg, ←mul_neg, complex.sin_eq_zero_iff, mul_comm] at hs, obtain ⟨k, hk⟩ := hs, rw [mul_eq_mul_right_iff, eq_false_intro (of_real_ne_zero.mpr pi_pos.ne'), or_false, - neg_eq_iff_neg_eq] at hk, - rw ←hk, + neg_eq_iff_eq_neg] at hk, + rw hk, cases k, { rw [int.cast_of_nat, complex.Gamma_neg_nat_eq_zero, zero_mul] }, { rw [int.cast_neg_succ_of_nat, neg_neg, nat.cast_add, nat.cast_one, add_comm, sub_add_cancel', diff --git a/src/analysis/special_functions/log/basic.lean b/src/analysis/special_functions/log/basic.lean index a02cc7d13b992..d70fb1ebd8237 100644 --- a/src/analysis/special_functions/log/basic.lean +++ b/src/analysis/special_functions/log/basic.lean @@ -177,9 +177,9 @@ begin split, { intros h, rcases lt_trichotomy x 0 with x_lt_zero | rfl | x_gt_zero, - { refine or.inr (or.inr (eq_neg_iff_eq_neg.mp _)), + { refine or.inr (or.inr (neg_eq_iff_eq_neg.mp _)), rw [←log_neg_eq_log x] at h, - exact (eq_one_of_pos_of_log_eq_zero (neg_pos.mpr x_lt_zero) h).symm, }, + exact eq_one_of_pos_of_log_eq_zero (neg_pos.mpr x_lt_zero) h, }, { exact or.inl rfl }, { exact or.inr (or.inl (eq_one_of_pos_of_log_eq_zero x_gt_zero h)), }, }, { rintro (rfl|rfl|rfl); simp only [log_one, log_zero, log_neg_eq_log], } diff --git a/src/analysis/special_functions/trigonometric/angle.lean b/src/analysis/special_functions/trigonometric/angle.lean index d4e36aee7c0f4..95c8108478a21 100644 --- a/src/analysis/special_functions/trigonometric/angle.lean +++ b/src/analysis/special_functions/trigonometric/angle.lean @@ -810,8 +810,8 @@ begin rcases hr with (hr|hr), { exact to_real_injective hr }, { by_cases h : θ = π, - { rw [h, to_real_pi, eq_neg_iff_eq_neg] at hr, - exact false.elim ((neg_pi_lt_to_real ψ).ne hr.symm) }, + { rw [h, to_real_pi, ← neg_eq_iff_eq_neg] at hr, + exact false.elim ((neg_pi_lt_to_real ψ).ne hr) }, { by_cases h' : ψ = π, { rw [h', to_real_pi] at hr, exact false.elim ((neg_pi_lt_to_real θ).ne hr.symm) }, diff --git a/src/data/int/basic.lean b/src/data/int/basic.lean index 433b7cb209ace..1554d6cc2650e 100644 --- a/src/data/int/basic.lean +++ b/src/data/int/basic.lean @@ -147,7 +147,7 @@ theorem succ_neg_succ (a : ℤ) : succ (-succ a) = -a := by rw [neg_succ, succ_pred] theorem neg_pred (a : ℤ) : -pred a = succ (-a) := -by rw [eq_neg_of_eq_neg (neg_succ (-a)).symm, neg_neg] +by rw [neg_eq_iff_eq_neg.mp (neg_succ (-a)), neg_neg] theorem pred_neg_pred (a : ℤ) : pred (-pred a) = -a := by rw [neg_pred, pred_succ] @@ -337,9 +337,8 @@ end theorem mod_add_div_aux (m n : ℕ) : (n - (m % n + 1) - (n * (m / n) + n) : ℤ) = -[1+ m] := begin - rw [← sub_sub, neg_succ_of_nat_coe, sub_sub (n:ℤ)], - apply eq_neg_of_eq_neg, - rw [neg_sub, sub_sub_self, add_right_comm], + rw [← sub_sub, neg_succ_of_nat_coe, sub_sub (n:ℤ), eq_comm, neg_eq_iff_eq_neg, + neg_sub, sub_sub_self, add_right_comm], exact @congr_arg ℕ ℤ _ _ (λi, (i + 1 : ℤ)) (nat.mod_add_div _ _).symm end diff --git a/src/data/num/lemmas.lean b/src/data/num/lemmas.lean index 949911b3da629..bfbfe1449518f 100644 --- a/src/data/num/lemmas.lean +++ b/src/data/num/lemmas.lean @@ -991,9 +991,9 @@ variables {α : Type*} have (↑b + -↑a : α) = -↑a + ↑b, by rw [← pos_num.cast_to_int a, ← pos_num.cast_to_int b, ← int.cast_neg, ← int.cast_add (-a)]; simp [add_comm], (pos_num.cast_sub' _ _).trans $ (sub_eq_add_neg _ _).trans this -| (neg a) (neg b) := show -(↑(a + b) : α) = -a + -b, by rw [ - pos_num.cast_add, neg_eq_iff_neg_eq, neg_add_rev, neg_neg, neg_neg, - ← pos_num.cast_to_int a, ← pos_num.cast_to_int b, ← int.cast_add]; simp [add_comm] +| (neg a) (neg b) := show -(↑(a + b) : α) = -a + -b, by rw [ + pos_num.cast_add, neg_eq_iff_eq_neg, neg_add_rev, neg_neg, neg_neg, + ← pos_num.cast_to_int a, ← pos_num.cast_to_int b, ← int.cast_add, ← int.cast_add, add_comm] @[simp] theorem cast_succ [add_group_with_one α] (n) : ((succ n : znum) : α) = n + 1 := by rw [← add_one, cast_add, cast_one] diff --git a/src/data/polynomial/mirror.lean b/src/data/polynomial/mirror.lean index e73d047aa1a96..2c39b8442a113 100644 --- a/src/data/polynomial/mirror.lean +++ b/src/data/polynomial/mirror.lean @@ -225,9 +225,9 @@ begin have hk := h2 k key, rcases hk with hk | hk | hk | hk, { exact or.inr (h3 h h_dvd_f (by rwa ← hk)) }, - { exact or.inr (h3 h h_dvd_f (by rwa [eq_neg_iff_eq_neg.mp hk, mirror_neg, dvd_neg])) }, + { exact or.inr (h3 h h_dvd_f (by rwa [← neg_eq_iff_eq_neg.mpr hk, mirror_neg, dvd_neg])) }, { exact or.inl (h3 g g_dvd_f (by rwa ← hk)) }, - { exact or.inl (h3 g g_dvd_f (by rwa [eq_neg_iff_eq_neg.mp hk, dvd_neg])) } }, + { exact or.inl (h3 g g_dvd_f (by rwa [← neg_eq_iff_eq_neg.mpr hk, dvd_neg])) } }, end end comm_ring diff --git a/src/data/real/conjugate_exponents.lean b/src/data/real/conjugate_exponents.lean index ad7f4833e48f3..11931ee0b8eba 100644 --- a/src/data/real/conjugate_exponents.lean +++ b/src/data/real/conjugate_exponents.lean @@ -68,8 +68,8 @@ ne_of_gt (h.one_div_pos) lemma conj_eq : q = p/(p-1) := begin have := h.inv_add_inv_conj, - rw [← eq_sub_iff_add_eq', one_div, inv_eq_iff_inv_eq] at this, - field_simp [← this, h.ne_zero] + rw [← eq_sub_iff_add_eq', one_div, inv_eq_iff_eq_inv] at this, + field_simp [this, h.ne_zero] end lemma conjugate_eq : conjugate_exponent p = q := h.conj_eq.symm diff --git a/src/data/real/ereal.lean b/src/data/real/ereal.lean index 833e1e4c87bb4..0e7f0f8f7d112 100644 --- a/src/data/real/ereal.lean +++ b/src/data/real/ereal.lean @@ -554,13 +554,13 @@ instance : has_involutive_neg ereal := | (x : ℝ) := rfl @[simp] lemma neg_eq_top_iff {x : ereal} : - x = ⊤ ↔ x = ⊥ := -by { rw neg_eq_iff_neg_eq, simp [eq_comm] } +neg_eq_iff_eq_neg @[simp] lemma neg_eq_bot_iff {x : ereal} : - x = ⊥ ↔ x = ⊤ := -by { rw neg_eq_iff_neg_eq, simp [eq_comm] } +neg_eq_iff_eq_neg @[simp] lemma neg_eq_zero_iff {x : ereal} : - x = 0 ↔ x = 0 := -by { rw neg_eq_iff_neg_eq, simp [eq_comm] } +by rw [neg_eq_iff_eq_neg, neg_zero] /-- if `-a ≤ b` then `-b ≤ a` on `ereal`. -/ protected theorem neg_le_of_neg_le {a b : ereal} (h : -a ≤ b) : -b ≤ a := diff --git a/src/data/real/golden_ratio.lean b/src/data/real/golden_ratio.lean index 8e9c3be8a18df..4c0ef8fec16ca 100644 --- a/src/data/real/golden_ratio.lean +++ b/src/data/real/golden_ratio.lean @@ -45,7 +45,7 @@ end /-- The opposite of the golden ratio is the inverse of its conjugate. -/ lemma inv_gold_conj : ψ⁻¹ = -φ := begin - rw [inv_eq_iff_inv_eq, ← neg_inv, neg_eq_iff_neg_eq], + rw [inv_eq_iff_eq_inv, ← neg_inv, ← neg_eq_iff_eq_neg], exact inv_gold.symm, end diff --git a/src/data/set/pointwise/interval.lean b/src/data/set/pointwise/interval.lean index b463eabdd2e8d..723d973f360d3 100644 --- a/src/data/set/pointwise/interval.lean +++ b/src/data/set/pointwise/interval.lean @@ -510,7 +510,7 @@ begin end lemma inv_Ioi {a : α} (ha : 0 < a) : (Ioi a)⁻¹ = Ioo 0 a⁻¹ := -by rw [inv_eq_iff_inv_eq, inv_Ioo_0_left (inv_pos.2 ha), inv_inv] +by rw [inv_eq_iff_eq_inv, inv_Ioo_0_left (inv_pos.2 ha), inv_inv] lemma image_const_mul_Ioi_zero {k : Type*} [linear_ordered_field k] {x : k} (hx : 0 < x) : diff --git a/src/field_theory/finite/basic.lean b/src/field_theory/finite/basic.lean index d75cbf1233591..1df22603ad038 100644 --- a/src/field_theory/finite/basic.lean +++ b/src/field_theory/finite/basic.lean @@ -101,7 +101,7 @@ begin have : (∏ x in (@univ Kˣ _).erase (-1), x) = 1, from prod_involution (λ x _, x⁻¹) (by simp) (λ a, by simp [units.inv_eq_self_iff] {contextual := tt}) - (λ a, by simp [@inv_eq_iff_inv_eq _ _ a, eq_comm]) + (λ a, by simp [@inv_eq_iff_eq_inv _ _ a]) (by simp), rw [← insert_erase (mem_univ (-1 : Kˣ)), prod_insert (not_mem_erase _ _), this, mul_one] diff --git a/src/geometry/euclidean/angle/oriented/basic.lean b/src/geometry/euclidean/angle/oriented/basic.lean index c4d739c33fcc8..4a79e60cf12f2 100644 --- a/src/geometry/euclidean/angle/oriented/basic.lean +++ b/src/geometry/euclidean/angle/oriented/basic.lean @@ -388,7 +388,7 @@ end /-- The oriented angle between two vectors is `π` if and only if the angle with the vectors swapped is `π`. -/ lemma oangle_eq_pi_iff_oangle_rev_eq_pi {x y : V} : o.oangle x y = π ↔ o.oangle y x = π := -by rw [oangle_rev, neg_eq_iff_neg_eq, eq_comm, real.angle.neg_coe_pi] +by rw [oangle_rev, neg_eq_iff_eq_neg, real.angle.neg_coe_pi] /-- The oriented angle between two vectors is `π` if and only they are nonzero and the first is on the same ray as the negation of the second. -/ @@ -700,7 +700,7 @@ begin by_cases hy : y = 0, { exfalso, simpa [hy] using h }, refine (o.oangle_eq_angle_or_eq_neg_angle hx hy).resolve_right _, intro hxy, - rw [hxy, real.angle.sign_neg, neg_eq_iff_neg_eq, eq_comm, ←sign_type.neg_iff, ←not_le] at h, + rw [hxy, real.angle.sign_neg, neg_eq_iff_eq_neg, ←sign_type.neg_iff, ←not_le] at h, exact h (real.angle.sign_coe_nonneg_of_nonneg_of_le_pi (inner_product_geometry.angle_nonneg _ _) (inner_product_geometry.angle_le_pi _ _)) end diff --git a/src/geometry/euclidean/angle/oriented/rotation.lean b/src/geometry/euclidean/angle/oriented/rotation.lean index 588502241a5e2..d0b86900c8488 100644 --- a/src/geometry/euclidean/angle/oriented/rotation.lean +++ b/src/geometry/euclidean/angle/oriented/rotation.lean @@ -183,7 +183,7 @@ by rw [neg_rotation, ←real.angle.coe_add, neg_div, ←sub_eq_add_neg, sub_half /-- Negating a rotation by π / 2 is equivalent to rotation by -π / 2. -/ lemma neg_rotation_pi_div_two (x : V) : -o.rotation (π / 2 : ℝ) x = o.rotation (-π / 2 : ℝ) x := -neg_eq_iff_neg_eq.1 $ o.neg_rotation_neg_pi_div_two _ +(neg_eq_iff_eq_neg.mp $ o.neg_rotation_neg_pi_div_two _).symm /-- Rotating the first of two vectors by `θ` scales their Kahler form by `cos (-θ) + sin (-θ) * I`. -/ diff --git a/src/linear_algebra/affine_space/midpoint.lean b/src/linear_algebra/affine_space/midpoint.lean index 167ffa4d2784f..727d674d2b8e0 100644 --- a/src/linear_algebra/affine_space/midpoint.lean +++ b/src/linear_algebra/affine_space/midpoint.lean @@ -133,7 +133,7 @@ by rw [eq_comm, midpoint_eq_right_iff] lemma midpoint_eq_midpoint_iff_vsub_eq_vsub {x x' y y' : P} : midpoint R x y = midpoint R x' y' ↔ x -ᵥ x' = y' -ᵥ y := by rw [← @vsub_eq_zero_iff_eq V, midpoint_vsub_midpoint, midpoint_eq_iff, point_reflection_apply, - vsub_eq_sub, zero_sub, vadd_eq_add, add_zero, neg_eq_iff_neg_eq, neg_vsub_eq_vsub_rev, eq_comm] + vsub_eq_sub, zero_sub, vadd_eq_add, add_zero, neg_eq_iff_eq_neg, neg_vsub_eq_vsub_rev] lemma midpoint_eq_iff' {x y z : P} : midpoint R x y = z ↔ equiv.point_reflection z x = y := midpoint_eq_iff diff --git a/src/linear_algebra/clifford_algebra/even_equiv.lean b/src/linear_algebra/clifford_algebra/even_equiv.lean index c16e72cfe02b9..caf1ccefbe519 100644 --- a/src/linear_algebra/clifford_algebra/even_equiv.lean +++ b/src/linear_algebra/clifford_algebra/even_equiv.lean @@ -71,8 +71,8 @@ end lemma neg_v_mul_e0 (m : M) : -(v Q m * e0 Q) = e0 Q * v Q m := begin - rw neg_eq_iff_neg_eq, - exact neg_e0_mul_v _ m + rw neg_eq_iff_eq_neg, + exact (neg_e0_mul_v _ m).symm end @[simp] lemma e0_mul_v_mul_e0 (m : M) : e0 Q * v Q m * e0 Q = v Q m := diff --git a/src/linear_algebra/dimension.lean b/src/linear_algebra/dimension.lean index 3b72dc742207d..6ccf5e71b30f0 100644 --- a/src/linear_algebra/dimension.lean +++ b/src/linear_algebra/dimension.lean @@ -390,7 +390,7 @@ begin -- by expressing the `v i` in the basis `b`, and using that the `v i` have no `b b'` term. have l₀ : l none = 0, { rw ←eq_neg_iff_add_eq_zero at z, - replace z := eq_neg_of_eq_neg z, + replace z := neg_eq_iff_eq_neg.mpr z, apply_fun (λ x, b.repr x b') at z, simp only [repr_self, linear_equiv.map_smul, mul_one, finsupp.single_eq_same, pi.neg_apply, finsupp.smul_single', linear_equiv.map_neg, finsupp.coe_neg] at z, diff --git a/src/measure_theory/integral/divergence_theorem.lean b/src/measure_theory/integral/divergence_theorem.lean index 1e6b317345a3c..17ecec3e87575 100644 --- a/src/measure_theory/integral/divergence_theorem.lean +++ b/src/measure_theory/integral/divergence_theorem.lean @@ -407,7 +407,7 @@ begin { simp only [uIcc_of_le hab, min_eq_left hab, max_eq_right hab] at *, exact integral_eq_of_has_deriv_within_at_off_countable_of_le f f' hab hs Hc Hd Hi }, { simp only [uIcc_of_ge hab, min_eq_right hab, max_eq_left hab] at *, - rw [interval_integral.integral_symm, neg_eq_iff_neg_eq, neg_sub, eq_comm], + rw [interval_integral.integral_symm, neg_eq_iff_eq_neg, neg_sub], exact integral_eq_of_has_deriv_within_at_off_countable_of_le f f' hab hs Hc Hd Hi.symm } end diff --git a/src/number_theory/bernoulli.lean b/src/number_theory/bernoulli.lean index a88c49a72d98b..b77fb4c739448 100644 --- a/src/number_theory/bernoulli.lean +++ b/src/number_theory/bernoulli.lean @@ -166,9 +166,9 @@ begin { simpa using h 1 } }, have h : B * (exp ℚ - 1) = X * exp ℚ, { simpa [bernoulli'_power_series] using bernoulli'_power_series_mul_exp_sub_one ℚ }, - rw [sub_mul, h, mul_sub X, sub_right_inj, ← neg_sub, mul_neg, neg_eq_iff_neg_eq], + rw [sub_mul, h, mul_sub X, sub_right_inj, ← neg_sub, mul_neg, neg_eq_iff_eq_neg], suffices : eval_neg_hom (B * (exp ℚ - 1)) * exp ℚ = eval_neg_hom (X * exp ℚ) * exp ℚ, - { simpa [mul_assoc, sub_mul, mul_comm (eval_neg_hom (exp ℚ)), exp_mul_exp_neg_eq_one, eq_comm] }, + { simpa [mul_assoc, sub_mul, mul_comm (eval_neg_hom (exp ℚ)), exp_mul_exp_neg_eq_one] }, congr', end diff --git a/src/number_theory/legendre_symbol/quadratic_reciprocity.lean b/src/number_theory/legendre_symbol/quadratic_reciprocity.lean index 93215318f6248..af92628a7b20c 100644 --- a/src/number_theory/legendre_symbol/quadratic_reciprocity.lean +++ b/src/number_theory/legendre_symbol/quadratic_reciprocity.lean @@ -235,7 +235,7 @@ end lemma mod_four_ne_three_of_sq_eq_neg_sq {x y : zmod p} (hx : x ≠ 0) (hxy : x ^ 2 = - y ^ 2) : p % 4 ≠ 3 := -mod_four_ne_three_of_sq_eq_neg_sq' hx (eq_neg_iff_eq_neg.1 hxy) +mod_four_ne_three_of_sq_eq_neg_sq' hx (neg_eq_iff_eq_neg.mpr hxy).symm end zmod diff --git a/src/number_theory/modular.lean b/src/number_theory/modular.lean index f998cd355aa84..901c51743781c 100644 --- a/src/number_theory/modular.lean +++ b/src/number_theory/modular.lean @@ -505,7 +505,7 @@ begin linarith, }, have hn : ↑ₘg 1 0 ≠ -1, { intros hc, - replace hc : ↑ₘ(-g) 1 0 = 1, { simp [eq_neg_of_eq_neg hc], }, + replace hc : ↑ₘ(-g) 1 0 = 1, { simp [← neg_eq_iff_eq_neg.mpr hc], }, replace hg : (-g) • z ∈ 𝒟ᵒ := (SL_neg_smul g z).symm ▸ hg, exact hp hg hc, }, specialize hp hg, diff --git a/src/probability/martingale/borel_cantelli.lean b/src/probability/martingale/borel_cantelli.lean index 363bfe0a98eee..fb9014347322b 100644 --- a/src/probability/martingale/borel_cantelli.lean +++ b/src/probability/martingale/borel_cantelli.lean @@ -256,7 +256,7 @@ begin refine neg_le.2 (hc _ _), simpa only [pi.neg_apply, set.mem_range, neg_inj] }, { rw mem_lower_bounds at hc, - simp_rw [set.mem_range, pi.neg_apply, neg_eq_iff_neg_eq, eq_comm] at hω, + simp_rw [set.mem_range, pi.neg_apply, neg_eq_iff_eq_neg] at hω, refine le_neg.1 (hc _ _), simpa only [set.mem_range] } end diff --git a/src/ring_theory/algebraic.lean b/src/ring_theory/algebraic.lean index d91e2bb41c483..603178ee4de13 100644 --- a/src/ring_theory/algebraic.lean +++ b/src/ring_theory/algebraic.lean @@ -299,7 +299,7 @@ lemma inv_eq_of_aeval_div_X_ne_zero {x : L} {p : K[X]} (aeval_ne : aeval x (div_X p) ≠ 0) : x⁻¹ = aeval x (div_X p) / (aeval x p - algebra_map _ _ (p.coeff 0)) := begin - rw [inv_eq_iff_inv_eq, inv_div, div_eq_iff, sub_eq_iff_eq_add, mul_comm], + rw [inv_eq_iff_eq_inv, inv_div, eq_comm, div_eq_iff, sub_eq_iff_eq_add, mul_comm], conv_lhs { rw ← div_X_mul_X_add p }, rw [alg_hom.map_add, alg_hom.map_mul, aeval_X, aeval_C], exact aeval_ne diff --git a/src/ring_theory/int/basic.lean b/src/ring_theory/int/basic.lean index cd86ca03afeb9..76f1f774b03a0 100644 --- a/src/ring_theory/int/basic.lean +++ b/src/ring_theory/int/basic.lean @@ -148,7 +148,7 @@ lemma exists_unit_of_abs (a : ℤ) : ∃ (u : ℤ) (h : is_unit u), (int.nat_abs begin cases (nat_abs_eq a) with h, { use [1, is_unit_one], rw [← h, one_mul], }, - { use [-1, is_unit_one.neg], rw [ ← neg_eq_iff_neg_eq.mp (eq.symm h)], + { use [-1, is_unit_one.neg], rw [← neg_eq_iff_eq_neg.mpr h], simp only [neg_mul, one_mul] } end diff --git a/src/ring_theory/valuation/basic.lean b/src/ring_theory/valuation/basic.lean index b3859e9586977..6acdabe47103b 100644 --- a/src/ring_theory/valuation/basic.lean +++ b/src/ring_theory/valuation/basic.lean @@ -420,14 +420,14 @@ begin by_contra h_1, cases ne_iff_lt_or_gt.1 h_1, { simpa [hh, lt_self_iff_false] using h.2 h_2 }, - { rw [← inv_one, eq_inv_iff_eq_inv, ← map_inv₀] at hh, - exact hh.le.not_lt (h.2 ((one_lt_val_iff v' hx).1 h_2)) } }, + { rw [← inv_one, ← inv_eq_iff_eq_inv, ← map_inv₀] at hh, + exact hh.not_lt (h.2 ((one_lt_val_iff v' hx).1 h_2)) } }, { intro hh, by_contra h_1, cases ne_iff_lt_or_gt.1 h_1, { simpa [hh, lt_self_iff_false] using h.1 h_2 }, - { rw [← inv_one, eq_inv_iff_eq_inv, ← map_inv₀] at hh, - exact hh.le.not_lt (h.1 ((one_lt_val_iff v hx).1 h_2)) } } } + { rw [← inv_one, ← inv_eq_iff_eq_inv, ← map_inv₀] at hh, + exact hh.not_lt (h.1 ((one_lt_val_iff v hx).1 h_2)) } } } end lemma is_equiv_iff_val_sub_one_lt_one diff --git a/src/ring_theory/valuation/valuation_subring.lean b/src/ring_theory/valuation/valuation_subring.lean index 49e106e62a016..faf017aa2555c 100644 --- a/src/ring_theory/valuation/valuation_subring.lean +++ b/src/ring_theory/valuation/valuation_subring.lean @@ -545,7 +545,7 @@ begin { intros h x hx, by_cases h_1 : x = 0, { simp only [h_1, zero_mem] }, by_cases h_2 : x⁻¹ + 1 = 0, - { rw [add_eq_zero_iff_eq_neg, inv_eq_iff_inv_eq, inv_neg, inv_one] at h_2, + { rw [add_eq_zero_iff_eq_neg, inv_eq_iff_eq_inv, inv_neg, inv_one] at h_2, simpa only [h_2] using B.neg_mem _ B.one_mem }, { rw [← valuation_le_one_iff, ← not_lt, valuation.one_lt_val_iff _ h_1, ← add_sub_cancel x⁻¹, ← units.coe_mk0 h_2, ← mem_principal_unit_group_iff] at hx ⊢, diff --git a/src/ring_theory/witt_vector/frobenius.lean b/src/ring_theory/witt_vector/frobenius.lean index c2d56ed005ce4..a5b97d84730a9 100644 --- a/src/ring_theory/witt_vector/frobenius.lean +++ b/src/ring_theory/witt_vector/frobenius.lean @@ -155,7 +155,7 @@ begin add_mul, add_mul, mul_right_comm, mul_right_comm (C (↑p ^ (n + 1))), ←C_mul, ←C_mul, pow_succ, mul_assoc ↑p (↑p ^ n), h1, mul_one, C_1, one_mul, add_comm _ (X n ^ p), add_assoc, ←add_sub, add_right_inj, frobenius_poly_aux_eq, ring_hom.map_sub, map_X, mul_sub, sub_eq_add_neg, - add_comm _ (C ↑p * X (n + 1)), ←add_sub, add_right_inj, neg_eq_iff_neg_eq, neg_sub], + add_comm _ (C ↑p * X (n + 1)), ←add_sub, add_right_inj, neg_eq_iff_eq_neg, neg_sub, eq_comm], simp only [ring_hom.map_sum, mul_sum, sum_mul, ←sum_sub_distrib], apply sum_congr rfl, intros i hi, diff --git a/src/set_theory/game/pgame.lean b/src/set_theory/game/pgame.lean index 3c9a085d0633f..15e2be320f042 100644 --- a/src/set_theory/game/pgame.lean +++ b/src/set_theory/game/pgame.lean @@ -844,7 +844,7 @@ theorem is_option_neg {x y : pgame} : is_option x (-y) ↔ is_option (-x) y := begin rw [is_option_iff, is_option_iff, or_comm], cases y, apply or_congr; - { apply exists_congr, intro, rw ← neg_eq_iff_neg_eq, exact eq_comm }, + { apply exists_congr, intro, rw neg_eq_iff_eq_neg, refl }, end @[simp] theorem is_option_neg_neg {x y : pgame} : is_option (-x) (-y) ↔ is_option x y := From bd1fc183335ea95a9519a1630bcf901fe9326d83 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 13 Mar 2023 22:49:50 +0000 Subject: [PATCH 0621/1291] refactor(linear_algebra/affine_space/affine_{map,equiv}): add fun_like instances (#18575) Going all the way and defining a new `affine_map_class` class can wait till after the port; but adding `fun_like` makes the port easier. This has to reorder a few declarations in `affine_equiv.lean`. The only new declarations are the new instances. --- .../affine_space/affine_equiv.lean | 55 +++++++++++-------- .../affine_space/affine_map.lean | 35 +++++++----- .../algebra/continuous_affine_map.lean | 40 +++++++------- 3 files changed, 71 insertions(+), 59 deletions(-) diff --git a/src/linear_algebra/affine_space/affine_equiv.lean b/src/linear_algebra/affine_space/affine_equiv.lean index 42559784f88d9..766c01c3a433b 100644 --- a/src/linear_algebra/affine_space/affine_equiv.lean +++ b/src/linear_algebra/affine_space/affine_equiv.lean @@ -59,36 +59,15 @@ namespace affine_equiv include V₁ V₂ -instance : has_coe_to_fun (P₁ ≃ᵃ[k] P₂) (λ _, P₁ → P₂) := ⟨λ e, e.to_fun⟩ - -instance : has_coe (P₁ ≃ᵃ[k] P₂) (P₁ ≃ P₂) := ⟨affine_equiv.to_equiv⟩ - -variables {k P₁} - -@[simp] lemma map_vadd (e : P₁ ≃ᵃ[k] P₂) (p : P₁) (v : V₁) : e (v +ᵥ p) = e.linear v +ᵥ e p := -e.map_vadd' p v - -@[simp] lemma coe_to_equiv (e : P₁ ≃ᵃ[k] P₂) : ⇑e.to_equiv = e := rfl - /-- Reinterpret an `affine_equiv` as an `affine_map`. -/ -def to_affine_map (e : P₁ ≃ᵃ[k] P₂) : P₁ →ᵃ[k] P₂ := { to_fun := e, .. e } - -instance : has_coe (P₁ ≃ᵃ[k] P₂) (P₁ →ᵃ[k] P₂) := ⟨to_affine_map⟩ - -@[simp] lemma coe_to_affine_map (e : P₁ ≃ᵃ[k] P₂) : - (e.to_affine_map : P₁ → P₂) = (e : P₁ → P₂) := -rfl +def to_affine_map (e : P₁ ≃ᵃ[k] P₂) : P₁ →ᵃ[k] P₂ := { .. e } @[simp] lemma to_affine_map_mk (f : P₁ ≃ P₂) (f' : V₁ ≃ₗ[k] V₂) (h) : to_affine_map (mk f f' h) = ⟨f, f', h⟩ := rfl -@[norm_cast, simp] lemma coe_coe (e : P₁ ≃ᵃ[k] P₂) : ((e : P₁ →ᵃ[k] P₂) : P₁ → P₂) = e := rfl - @[simp] lemma linear_to_affine_map (e : P₁ ≃ᵃ[k] P₂) : e.to_affine_map.linear = e.linear := rfl -@[simp] lemma coe_linear (e : P₁ ≃ᵃ[k] P₂) : (e : P₁ →ᵃ[k] P₂).linear = e.linear := rfl - lemma to_affine_map_injective : injective (to_affine_map : (P₁ ≃ᵃ[k] P₂) → (P₁ →ᵃ[k] P₂)) := begin rintros ⟨e, el, h⟩ ⟨e', el', h'⟩ H, @@ -101,11 +80,39 @@ end e.to_affine_map = e'.to_affine_map ↔ e = e' := to_affine_map_injective.eq_iff +instance equiv_like : equiv_like (P₁ ≃ᵃ[k] P₂) P₁ P₂ := +{ coe := λ f, f.to_fun, + inv := λ f, f.inv_fun, + left_inv := λ f, f.left_inv, + right_inv := λ f, f.right_inv, + coe_injective' := λ f g h _, to_affine_map_injective (fun_like.coe_injective h) } + +instance : has_coe_to_fun (P₁ ≃ᵃ[k] P₂) (λ _, P₁ → P₂) := fun_like.has_coe_to_fun + +instance : has_coe (P₁ ≃ᵃ[k] P₂) (P₁ ≃ P₂) := ⟨affine_equiv.to_equiv⟩ + +variables {k P₁} + +@[simp] lemma map_vadd (e : P₁ ≃ᵃ[k] P₂) (p : P₁) (v : V₁) : e (v +ᵥ p) = e.linear v +ᵥ e p := +e.map_vadd' p v + +@[simp] lemma coe_to_equiv (e : P₁ ≃ᵃ[k] P₂) : ⇑e.to_equiv = e := rfl + +instance : has_coe (P₁ ≃ᵃ[k] P₂) (P₁ →ᵃ[k] P₂) := ⟨to_affine_map⟩ + +@[simp] lemma coe_to_affine_map (e : P₁ ≃ᵃ[k] P₂) : + (e.to_affine_map : P₁ → P₂) = (e : P₁ → P₂) := +rfl + +@[norm_cast, simp] lemma coe_coe (e : P₁ ≃ᵃ[k] P₂) : ((e : P₁ →ᵃ[k] P₂) : P₁ → P₂) = e := rfl + +@[simp] lemma coe_linear (e : P₁ ≃ᵃ[k] P₂) : (e : P₁ →ᵃ[k] P₂).linear = e.linear := rfl + @[ext] lemma ext {e e' : P₁ ≃ᵃ[k] P₂} (h : ∀ x, e x = e' x) : e = e' := -to_affine_map_injective $ affine_map.ext h +fun_like.ext _ _ h lemma coe_fn_injective : @injective (P₁ ≃ᵃ[k] P₂) (P₁ → P₂) coe_fn := -λ e e' H, ext $ congr_fun H +fun_like.coe_injective @[simp, norm_cast] lemma coe_fn_inj {e e' : P₁ ≃ᵃ[k] P₂} : (e : P₁ → P₂) = e' ↔ e = e' := coe_fn_injective.eq_iff diff --git a/src/linear_algebra/affine_space/affine_map.lean b/src/linear_algebra/affine_space/affine_map.lean index 2e52c959038ef..3714083f07b4f 100644 --- a/src/linear_algebra/affine_space/affine_map.lean +++ b/src/linear_algebra/affine_space/affine_map.lean @@ -59,11 +59,24 @@ structure affine_map (k : Type*) {V1 : Type*} (P1 : Type*) {V2 : Type*} (P2 : Ty notation P1 ` →ᵃ[`:25 k:25 `] `:0 P2:0 := affine_map k P1 P2 -instance (k : Type*) {V1 : Type*} (P1 : Type*) {V2 : Type*} (P2 : Type*) - [ring k] - [add_comm_group V1] [module k V1] [affine_space V1 P1] - [add_comm_group V2] [module k V2] [affine_space V2 P2]: - has_coe_to_fun (P1 →ᵃ[k] P2) (λ _, P1 → P2) := ⟨affine_map.to_fun⟩ +instance affine_map.fun_like (k : Type*) {V1 : Type*} (P1 : Type*) {V2 : Type*} (P2 : Type*) + [ring k] + [add_comm_group V1] [module k V1] [affine_space V1 P1] + [add_comm_group V2] [module k V2] [affine_space V2 P2]: + fun_like (P1 →ᵃ[k] P2) P1 (λ _, P2) := +{ coe := affine_map.to_fun, + coe_injective' := λ ⟨f, f_linear, f_add⟩ ⟨g, g_linear, g_add⟩ (h : f = g), begin + cases (add_torsor.nonempty : nonempty P1) with p, + congr' with v, + apply vadd_right_cancel (f p), + erw [← f_add, h, ← g_add] + end } + +instance affine_map.has_coe_to_fun (k : Type*) {V1 : Type*} (P1 : Type*) {V2 : Type*} (P2 : Type*) + [ring k] + [add_comm_group V1] [module k V1] [affine_space V1 P1] + [add_comm_group V2] [module k V2] [affine_space V2 P2] : + has_coe_to_fun (P1 →ᵃ[k] P2) (λ _, P1 → P2) := fun_like.has_coe_to_fun namespace linear_map @@ -115,20 +128,12 @@ by conv_rhs { rw [←vsub_vadd p1 p2, map_vadd, vadd_vsub] } /-- Two affine maps are equal if they coerce to the same function. -/ @[ext] lemma ext {f g : P1 →ᵃ[k] P2} (h : ∀ p, f p = g p) : f = g := -begin - rcases f with ⟨f, f_linear, f_add⟩, - rcases g with ⟨g, g_linear, g_add⟩, - obtain rfl : f = g := funext h, - congr' with v, - cases (add_torsor.nonempty : nonempty P1) with p, - apply vadd_right_cancel (f p), - erw [← f_add, ← g_add] -end +fun_like.ext _ _ h lemma ext_iff {f g : P1 →ᵃ[k] P2} : f = g ↔ ∀ p, f p = g p := ⟨λ h p, h ▸ rfl, ext⟩ lemma coe_fn_injective : @function.injective (P1 →ᵃ[k] P2) (P1 → P2) coe_fn := -λ f g H, ext $ congr_fun H +fun_like.coe_injective protected lemma congr_arg (f : P1 →ᵃ[k] P2) {x y : P1} (h : x = y) : f x = f y := congr_arg _ h diff --git a/src/topology/algebra/continuous_affine_map.lean b/src/topology/algebra/continuous_affine_map.lean index eb16136d168de..df8ea46d78521 100644 --- a/src/topology/algebra/continuous_affine_map.lean +++ b/src/topology/algebra/continuous_affine_map.lean @@ -45,30 +45,34 @@ variables [add_comm_group W] [module R W] [topological_space Q] [add_torsor W Q] include V W -/-- see Note [function coercion] -/ -instance : has_coe_to_fun (P →A[R] Q) (λ _, P → Q) := ⟨λ f, f.to_affine_map.to_fun⟩ +instance : has_coe (P →A[R] Q) (P →ᵃ[R] Q) := +⟨to_affine_map⟩ + +lemma to_affine_map_injective {f g : P →A[R] Q} (h : (f : P →ᵃ[R] Q) = (g : P →ᵃ[R] Q)) : f = g := +by { cases f, cases g, congr' } + +instance : continuous_map_class (P →A[R] Q) P Q := +{ coe := λ f, f.to_affine_map, + coe_injective' := λ f g h, to_affine_map_injective $ fun_like.coe_injective h, + map_continuous := cont } + +/-- Helper instance for when there's too many metavariables to apply +`fun_like.has_coe_to_fun` directly. -/ +instance : has_coe_to_fun (P →A[R] Q) (λ _, P → Q) := fun_like.has_coe_to_fun lemma to_fun_eq_coe (f : P →A[R] Q) : f.to_fun = ⇑f := rfl -lemma coe_injective : - @function.injective (P →A[R] Q) (P → Q) coe_fn := -begin - rintros ⟨⟨f, ⟨f', hf₁, hf₂⟩, hf₀⟩, hf₁⟩ ⟨⟨g, ⟨g', hg₁, hg₂⟩, hg₀⟩, hg₁⟩ h, - have : f = g ∧ f' = g', { simpa only using affine_map.coe_fn_injective h, }, - congr, - exacts [this.1, this.2], -end +lemma coe_injective : @function.injective (P →A[R] Q) (P → Q) coe_fn := +fun_like.coe_injective @[ext] lemma ext {f g : P →A[R] Q} (h : ∀ x, f x = g x) : f = g := -coe_injective $ funext h +fun_like.ext _ _ h lemma ext_iff {f g : P →A[R] Q} : f = g ↔ ∀ x, f x = g x := -⟨by { rintro rfl x, refl, }, ext⟩ +fun_like.ext_iff -lemma congr_fun {f g : P →A[R] Q} (h : f = g) (x : P) : f x = g x := h ▸ rfl - -instance : has_coe (P →A[R] Q) (P →ᵃ[R] Q) := -⟨to_affine_map⟩ +lemma congr_fun {f g : P →A[R] Q} (h : f = g) (x : P) : f x = g x := +fun_like.congr_fun h _ /-- Forgetting its algebraic properties, a continuous affine map is a continuous map. -/ def to_continuous_map (f : P →A[R] Q) : C(P, Q) := @@ -92,10 +96,6 @@ rfl ((f : C(P, Q)) : P → Q) = f := rfl -lemma to_affine_map_injective {f g : P →A[R] Q} - (h : (f : P →ᵃ[R] Q) = (g : P →ᵃ[R] Q)) : f = g := -by { ext a, exact affine_map.congr_fun h a, } - lemma to_continuous_map_injective {f g : P →A[R] Q} (h : (f : C(P, Q)) = (g : C(P, Q))) : f = g := by { ext a, exact continuous_map.congr_fun h a, } From 2fd0de23fbbf0f8a84dbb7bc2e7a2da45de67f8b Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 14 Mar 2023 04:08:23 +0000 Subject: [PATCH 0622/1291] chore(analysis/normed_space/pi_Lp): missing scalar tower instances (#18577) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Extracted from #6799. Note these have rather strong requirements on `𝕜` and `𝕜''` as those are the requirements for there to be a `smul` operation. --- src/analysis/normed_space/pi_Lp.lean | 16 +++++++++++++--- 1 file changed, 13 insertions(+), 3 deletions(-) diff --git a/src/analysis/normed_space/pi_Lp.lean b/src/analysis/normed_space/pi_Lp.lean index 881441f216f65..da14f060696de 100644 --- a/src/analysis/normed_space/pi_Lp.lean +++ b/src/analysis/normed_space/pi_Lp.lean @@ -75,7 +75,7 @@ instance (p : ℝ≥0∞) {ι : Type*} (α : ι → Type*) [Π i, inhabited (α namespace pi_Lp -variables (p : ℝ≥0∞) (𝕜 : Type*) {ι : Type*} (α : ι → Type*) (β : ι → Type*) +variables (p : ℝ≥0∞) (𝕜 𝕜' : Type*) {ι : Type*} (α : ι → Type*) (β : ι → Type*) /-- Canonical bijection between `pi_Lp p α` and the original Pi type. We introduce it to be able to compare the `L^p` and `L^∞` distances through it. -/ @@ -526,7 +526,7 @@ lemma edist_eq_of_L2 {β : ι → Type*} [Π i, seminormed_add_comm_group (β i) edist x y = (∑ i, edist (x i) (y i) ^ 2) ^ (1 / 2 : ℝ) := by simp [pi_Lp.edist_eq_sum] -variables [normed_field 𝕜] +variables [normed_field 𝕜] [normed_field 𝕜'] /-- The product of finitely many normed spaces is a normed space, with the `L^p` norm. -/ instance normed_space [Π i, seminormed_add_comm_group (β i)] @@ -546,6 +546,16 @@ instance normed_space [Π i, seminormed_add_comm_group (β i)] end, .. (pi.module ι β 𝕜) } +instance is_scalar_tower [Π i, seminormed_add_comm_group (β i)] + [has_smul 𝕜 𝕜'] [Π i, normed_space 𝕜 (β i)] [Π i, normed_space 𝕜' (β i)] + [Π i, is_scalar_tower 𝕜 𝕜' (β i)] : is_scalar_tower 𝕜 𝕜' (pi_Lp p β) := +pi.is_scalar_tower + +instance smul_comm_class [Π i, seminormed_add_comm_group (β i)] + [Π i, normed_space 𝕜 (β i)] [Π i, normed_space 𝕜' (β i)] + [Π i, smul_comm_class 𝕜 𝕜' (β i)] : smul_comm_class 𝕜 𝕜' (pi_Lp p β) := +pi.smul_comm_class + instance finite_dimensional [Π i, seminormed_add_comm_group (β i)] [Π i, normed_space 𝕜 (β i)] [I : ∀ i, finite_dimensional 𝕜 (β i)] : finite_dimensional 𝕜 (pi_Lp p β) := @@ -553,7 +563,7 @@ finite_dimensional.finite_dimensional_pi' _ _ /- Register simplification lemmas for the applications of `pi_Lp` elements, as the usual lemmas for Pi types will not trigger. -/ -variables {𝕜 p α} [Π i, seminormed_add_comm_group (β i)] [Π i, normed_space 𝕜 (β i)] (c : 𝕜) +variables {𝕜 𝕜' p α} [Π i, seminormed_add_comm_group (β i)] [Π i, normed_space 𝕜 (β i)] (c : 𝕜) variables (x y : pi_Lp p β) (x' y' : Π i, β i) (i : ι) @[simp] lemma zero_apply : (0 : pi_Lp p β) i = 0 := rfl From ceb887ddf3344dab425292e497fa2af91498437c Mon Sep 17 00:00:00 2001 From: Alex J Best Date: Tue, 14 Mar 2023 11:21:55 +0000 Subject: [PATCH 0623/1291] feat(algebra/char_p/basic): refactor proof of add_pow_char_of_commute to extract a statement true in all rings (#11364) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit If x and y commute, the pth power of their sum is the sum of the pth powers plus a multiple of p. This holds in any semiring and easily implies the additivity of frobenius in char p, but is occasionally useful in characteristic 0. Also make the fact that p is 0 in a char_p ring a simp lemma. From flt-regular Co-authored-by: Johan Commelin Co-authored-by: Yaël Dillies Co-authored-by: Eric Wieser --- src/algebra/char_p/basic.lean | 98 ++++++++++++++----- src/combinatorics/additive/behrend.lean | 4 +- src/data/enat/basic.lean | 3 +- src/data/nat/multiplicity.lean | 17 +++- src/ring_theory/multiplicity.lean | 4 +- .../polynomial/eisenstein/is_integral.lean | 2 +- 6 files changed, 95 insertions(+), 33 deletions(-) diff --git a/src/algebra/char_p/basic.lean b/src/algebra/char_p/basic.lean index c93798109be1c..2346c09251dad 100644 --- a/src/algebra/char_p/basic.lean +++ b/src/algebra/char_p/basic.lean @@ -3,11 +3,8 @@ Copyright (c) 2018 Kenny Lau. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau, Joey van Langen, Casper Putz -/ - -import algebra.hom.iterate import data.int.modeq -import data.nat.choose.dvd -import data.nat.choose.sum +import data.nat.multiplicity import group_theory.order_of_element import ring_theory.nilpotent @@ -17,7 +14,70 @@ import ring_theory.nilpotent universes u v -variables (R : Type u) +open finset +open_locale big_operators + +variables {R : Type*} + +namespace commute +variables [semiring R] {p : ℕ} {x y : R} + +protected lemma add_pow_prime_pow_eq (hp : p.prime) (h : commute x y) (n : ℕ) : + (x + y) ^ p ^ n = x ^ p ^ n + y ^ p ^ n + + p * ∑ k in Ioo 0 (p ^ n), x ^ k * y ^ (p ^ n - k) * ↑((p ^ n).choose k / p) := +begin + transitivity + x ^ p ^ n + y ^ p ^ n + ∑ k in Ioo 0 (p ^ n), x ^ k * y ^ (p ^ n - k) * (p ^ n).choose k, + { simp_rw [h.add_pow, ←nat.Ico_zero_eq_range, nat.Ico_succ_right, Icc_eq_cons_Ico (zero_le _), + finset.sum_cons, Ico_eq_cons_Ioo (pow_pos hp.pos _), finset.sum_cons, tsub_self, tsub_zero, + pow_zero, nat.choose_zero_right, nat.choose_self, nat.cast_one, mul_one, one_mul, + ←add_assoc] }, + { congr' 1, + simp_rw [finset.mul_sum, nat.cast_comm, mul_assoc _ _ (p : R), ←nat.cast_mul], + refine finset.sum_congr rfl (λ i hi, _), + rw mem_Ioo at hi, + rw nat.div_mul_cancel (hp.dvd_choose_pow hi.1.ne' hi.2.ne) }, +end + +protected lemma add_pow_prime_eq (hp : p.prime) (h : commute x y) : + (x + y) ^ p = x ^ p + y ^ p + + p * ∑ k in finset.Ioo 0 p, x ^ k * y ^ (p - k) * ↑(p.choose k / p) := +by simpa using h.add_pow_prime_pow_eq hp 1 + +protected lemma exists_add_pow_prime_pow_eq (hp : p.prime) (h : commute x y) (n : ℕ) : + ∃ r, (x + y) ^ p ^ n = x ^ p ^ n + y ^ p ^ n + p * r := +⟨_, h.add_pow_prime_pow_eq hp n⟩ + +protected lemma exists_add_pow_prime_eq (hp : p.prime) (h : commute x y) : + ∃ r, (x + y) ^ p = x ^ p + y ^ p + p * r := +⟨_, h.add_pow_prime_eq hp⟩ + +end commute + +section comm_semiring +variables [comm_semiring R] {p : ℕ} {x y : R} + +lemma add_pow_prime_pow_eq (hp : p.prime) (x y : R) (n : ℕ) : + (x + y) ^ p ^ n = x ^ p ^ n + y ^ p ^ n + + p * ∑ k in finset.Ioo 0 (p ^ n), x ^ k * y ^ (p ^ n - k) * ↑((p ^ n).choose k / p) := +(commute.all x y).add_pow_prime_pow_eq hp n + +lemma add_pow_prime_eq (hp : p.prime) (x y : R) : + (x + y) ^ p = x ^ p + y ^ p + + p * ∑ k in finset.Ioo 0 p, x ^ k * y ^ (p - k) * ↑(p.choose k / p) := +(commute.all x y).add_pow_prime_eq hp + +lemma exists_add_pow_prime_pow_eq (hp : p.prime) (x y : R) (n : ℕ) : + ∃ r, (x + y) ^ p ^ n = x ^ p ^ n + y ^ p ^ n + p * r := +(commute.all x y).exists_add_pow_prime_pow_eq hp n + +lemma exists_add_pow_prime_eq (hp : p.prime) (x y : R) : + ∃ r, (x + y) ^ p = x ^ p + y ^ p + p * r := +(commute.all x y).exists_add_pow_prime_eq hp + +end comm_semiring + +variables (R) /-- The generator of the kernel of the unique homomorphism ℕ → R for a semiring R. @@ -33,6 +93,7 @@ This example is formalized in `counterexamples/char_p_zero_ne_char_zero`. class char_p [add_monoid_with_one R] (p : ℕ) : Prop := (cast_eq_zero_iff [] : ∀ x:ℕ, (x:R) = 0 ↔ p ∣ x) +@[simp] theorem char_p.cast_eq_zero [add_monoid_with_one R] (p : ℕ) [char_p R p] : (p:R) = 0 := (char_p.cast_eq_zero_iff R p p).2 (dvd_refl p) @@ -133,32 +194,15 @@ by rw ring_char.spec end ring_char -theorem add_pow_char_of_commute [semiring R] {p : ℕ} [fact p.prime] +theorem add_pow_char_of_commute [semiring R] {p : ℕ} [hp : fact p.prime] [char_p R p] (x y : R) (h : commute x y) : (x + y)^p = x^p + y^p := -begin - rw [commute.add_pow h, finset.sum_range_succ_comm, tsub_self, pow_zero, nat.choose_self], - rw [nat.cast_one, mul_one, mul_one], congr' 1, - convert finset.sum_eq_single 0 _ _, - { simp only [mul_one, one_mul, nat.choose_zero_right, tsub_zero, nat.cast_one, pow_zero] }, - { intros b h1 h2, - suffices : (p.choose b : R) = 0, { rw this, simp }, - rw char_p.cast_eq_zero_iff R p, - exact nat.prime.dvd_choose_self (fact.out _) h2 (finset.mem_range.1 h1) }, - { intro h1, - contrapose! h1, - rw finset.mem_range, - exact nat.prime.pos (fact.out _) } -end +let ⟨r, hr⟩ := h.exists_add_pow_prime_eq hp.out in by simp [hr] -theorem add_pow_char_pow_of_commute [semiring R] {p : ℕ} [fact p.prime] - [char_p R p] {n : ℕ} (x y : R) (h : commute x y) : +theorem add_pow_char_pow_of_commute [semiring R] {p n : ℕ} [hp : fact p.prime] [char_p R p] + (x y : R) (h : commute x y) : (x + y) ^ (p ^ n) = x ^ (p ^ n) + y ^ (p ^ n) := -begin - induction n, { simp, }, - rw [pow_succ', pow_mul, pow_mul, pow_mul, n_ih], - apply add_pow_char_of_commute, apply commute.pow_pow h, -end +let ⟨r, hr⟩ := h.exists_add_pow_prime_pow_eq hp.out n in by simp [hr] theorem sub_pow_char_of_commute [ring R] {p : ℕ} [fact p.prime] [char_p R p] (x y : R) (h : commute x y) : diff --git a/src/combinatorics/additive/behrend.lean b/src/combinatorics/additive/behrend.lean index a2c1ac3cd379e..857a123b6c14e 100644 --- a/src/combinatorics/additive/behrend.lean +++ b/src/combinatorics/additive/behrend.lean @@ -42,7 +42,7 @@ integer points on that sphere and map them onto `ℕ` in a way that preserves ar Salem-Spencer, Behrend construction, arithmetic progression, sphere, strictly convex -/ -open finset nat real +open finset nat (hiding log) real open_locale big_operators pointwise namespace behrend @@ -432,7 +432,7 @@ begin rw one_le_cast, exact hN.trans' (by norm_num1) }, { rw [cast_pos, lt_ceil, cast_zero, real.sqrt_pos], - apply log_pos, + refine log_pos _, rw one_lt_cast, exact hN.trans_lt' (by norm_num1) }, apply le_sqrt_of_sq_le, diff --git a/src/data/enat/basic.lean b/src/data/enat/basic.lean index b344502039c9d..105dac551f4f0 100644 --- a/src/data/enat/basic.lean +++ b/src/data/enat/basic.lean @@ -34,7 +34,8 @@ instance : is_well_order ℕ∞ (<) := { } variables {m n : ℕ∞} -@[simp, norm_cast] lemma coe_zero : ((0 : ℕ) : ℕ∞) = 0 := rfl +-- eligible for `dsimp` +@[simp, nolint simp_nf, norm_cast] lemma coe_zero : ((0 : ℕ) : ℕ∞) = 0 := rfl @[simp, norm_cast] lemma coe_one : ((1 : ℕ) : ℕ∞) = 1 := rfl @[simp, norm_cast] lemma coe_add (m n : ℕ) : ↑(m + n) = (m + n : ℕ∞) := rfl @[simp, norm_cast] lemma coe_sub (m n : ℕ) : ↑(m - n) = (m - n : ℕ∞) := rfl diff --git a/src/data/nat/multiplicity.lean b/src/data/nat/multiplicity.lean index 0f1f39d09b6f7..68a56270c1855 100644 --- a/src/data/nat/multiplicity.lean +++ b/src/data/nat/multiplicity.lean @@ -203,7 +203,9 @@ begin exact dvd_mul_right _ _ end -lemma multiplicity_choose_prime_pow_add_multiplicity {p n k : ℕ} (hp : p.prime) (hkn : k ≤ p ^ n) +variables {p n k : ℕ} + +lemma multiplicity_choose_prime_pow_add_multiplicity (hp : p.prime) (hkn : k ≤ p ^ n) (hk0 : k ≠ 0) : multiplicity p (choose (p ^ n) k) + multiplicity p k = n := le_antisymm @@ -229,6 +231,19 @@ lemma multiplicity_choose_prime_pow {p n k : ℕ} (hp : p.prime) (hkn : k ≤ p ↑(n - (multiplicity p k).get (finite_nat_iff.2 ⟨hp.ne_one, hk0.bot_lt⟩)) := part_enat.eq_coe_sub_of_add_eq_coe $ multiplicity_choose_prime_pow_add_multiplicity hp hkn hk0 +lemma dvd_choose_pow (hp : prime p) (hk : k ≠ 0) (hkp : k ≠ p ^ n) : p ∣ (p ^ n).choose k := +begin + obtain hkp | hkp := hkp.symm.lt_or_lt, + { simp [choose_eq_zero_of_lt hkp] }, + refine multiplicity_ne_zero.1 (λ h, hkp.not_le $ nat.le_of_dvd hk.bot_lt _), + have H := hp.multiplicity_choose_prime_pow_add_multiplicity hkp.le hk, + rw [h, zero_add, eq_coe_iff] at H, + exact H.1, +end + +lemma dvd_choose_pow_iff (hp : prime p) : p ∣ (p ^ n).choose k ↔ k ≠ 0 ∧ k ≠ p ^ n := +by refine ⟨λ h, ⟨_, _⟩, λ h, dvd_choose_pow hp h.1 h.2⟩; rintro rfl; simpa [hp.ne_one] using h + end prime lemma multiplicity_two_factorial_lt : ∀ {n : ℕ} (h : n ≠ 0), multiplicity 2 n! < n := diff --git a/src/ring_theory/multiplicity.lean b/src/ring_theory/multiplicity.lean index fe9116d68cdc9..3a263db7e015b 100644 --- a/src/ring_theory/multiplicity.lean +++ b/src/ring_theory/multiplicity.lean @@ -145,9 +145,11 @@ end @[simp] lemma unit_left (a : α) (u : αˣ) : multiplicity (u : α) a = ⊤ := is_unit_left a u.is_unit -lemma multiplicity_eq_zero {a b : α} : multiplicity a b = 0 ↔ ¬a ∣ b := +lemma multiplicity_eq_zero {a b : α} : multiplicity a b = 0 ↔ ¬ a ∣ b := by { rw [← nat.cast_zero, eq_coe_iff], simp } +lemma multiplicity_ne_zero {a b : α} : multiplicity a b ≠ 0 ↔ a ∣ b := multiplicity_eq_zero.not_left + lemma eq_top_iff_not_finite {a b : α} : multiplicity a b = ⊤ ↔ ¬ finite a b := part.eq_none_iff' diff --git a/src/ring_theory/polynomial/eisenstein/is_integral.lean b/src/ring_theory/polynomial/eisenstein/is_integral.lean index 79c85cb016f1a..8c44b9246e35d 100644 --- a/src/ring_theory/polynomial/eisenstein/is_integral.lean +++ b/src/ring_theory/polynomial/eisenstein/is_integral.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Riccardo Brasca. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Riccardo Brasca -/ - +import data.nat.choose.dvd import ring_theory.integrally_closed import ring_theory.norm import ring_theory.polynomial.cyclotomic.basic From af8f9bc2aab1dd99a26f3dc982b7957f7762e7ba Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 14 Mar 2023 13:51:43 +0000 Subject: [PATCH 0624/1291] chore(analysis/inner_product_space/basic): make arguments explicit (#18579) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This applies the regex ```regex \{(x[^}]+)\}(?![^\n]*↔)(?=[^\n]* =) ``` and replaces it with `($1)` in appropriate places. The heuristic is "equalities should take their arguments explicitly". Underscores are then inserted as necessary to fix the build. --- src/analysis/convex/cone/basic.lean | 6 +- src/analysis/inner_product_space/basic.lean | 189 +++++++++--------- .../inner_product_space/lax_milgram.lean | 2 +- .../inner_product_space/projection.lean | 2 +- .../inner_product_space/symmetric.lean | 4 +- src/measure_theory/function/l2_space.lean | 2 +- 6 files changed, 103 insertions(+), 102 deletions(-) diff --git a/src/analysis/convex/cone/basic.lean b/src/analysis/convex/cone/basic.lean index 599fea0333854..a619278fbb11a 100644 --- a/src/analysis/convex/cone/basic.lean +++ b/src/analysis/convex/cone/basic.lean @@ -743,14 +743,14 @@ eq_top_iff.mpr $ λ x hy y, false.elim /-- Dual cone of the convex cone {0} is the total space. -/ @[simp] lemma inner_dual_cone_zero : (0 : set H).inner_dual_cone = ⊤ := -eq_top_iff.mpr $ λ x hy y (hy : y = 0), hy.symm ▸ inner_zero_left.ge +eq_top_iff.mpr $ λ x hy y (hy : y = 0), hy.symm ▸ (inner_zero_left _).ge /-- Dual cone of the total space is the convex cone {0}. -/ @[simp] lemma inner_dual_cone_univ : (univ : set H).inner_dual_cone = 0 := begin suffices : ∀ x : H, x ∈ (univ : set H).inner_dual_cone → x = 0, { apply set_like.coe_injective, - exact eq_singleton_iff_unique_mem.mpr ⟨λ x hx, inner_zero_right.ge, this⟩ }, + exact eq_singleton_iff_unique_mem.mpr ⟨λ x hx, (inner_zero_right _).ge, this⟩ }, exact λ x hx, by simpa [←real_inner_self_nonpos] using hx (-x) (mem_univ _), end @@ -874,7 +874,7 @@ begin calc 0 < ⟪b - z, b - z⟫_ℝ : lt_of_not_le ((iff.not real_inner_self_nonpos).2 hbz) ... = ⟪b - z, b - z⟫_ℝ + 0 : (add_zero _).symm ... ≤ ⟪b - z, b - z⟫_ℝ + ⟪b - z, z⟫_ℝ : add_le_add rfl.ge hinner₀ - ... = ⟪b - z, b - z + z⟫_ℝ : inner_add_right.symm + ... = ⟪b - z, b - z + z⟫_ℝ : (inner_add_right _ _ _).symm ... = ⟪b - z, b⟫_ℝ : by rw sub_add_cancel }, end diff --git a/src/analysis/inner_product_space/basic.lean b/src/analysis/inner_product_space/basic.lean index 9684da3de7d98..0684b1c478e6b 100644 --- a/src/analysis/inner_product_space/basic.lean +++ b/src/analysis/inner_product_space/basic.lean @@ -170,16 +170,16 @@ lemma inner_conj_sym (x y : F) : ⟪y, x⟫† = ⟪x, y⟫ := c.conj_sym x y lemma inner_self_nonneg {x : F} : 0 ≤ re ⟪x, x⟫ := c.nonneg_re _ -lemma inner_self_nonneg_im {x : F} : im ⟪x, x⟫ = 0 := +lemma inner_self_nonneg_im (x : F) : im ⟪x, x⟫ = 0 := by rw [← @of_real_inj 𝕜, im_eq_conj_sub]; simp [inner_conj_sym] -lemma inner_self_im_zero {x : F} : im ⟪x, x⟫ = 0 := -inner_self_nonneg_im +lemma inner_self_im_zero (x : F) : im ⟪x, x⟫ = 0 := +inner_self_nonneg_im _ -lemma inner_add_left {x y z : F} : ⟪x + y, z⟫ = ⟪x, z⟫ + ⟪y, z⟫ := +lemma inner_add_left (x y z : F) : ⟪x + y, z⟫ = ⟪x, z⟫ + ⟪y, z⟫ := c.add_left _ _ _ -lemma inner_add_right {x y z : F} : ⟪x, y + z⟫ = ⟪x, y⟫ + ⟪x, z⟫ := +lemma inner_add_right (x y z : F) : ⟪x, y + z⟫ = ⟪x, y⟫ + ⟪x, z⟫ := by rw [←inner_conj_sym, inner_add_left, ring_hom.map_add]; simp only [inner_conj_sym] lemma inner_norm_sq_eq_inner_self (x : F) : (norm_sqF x : 𝕜) = ⟪x, x⟫ := @@ -188,54 +188,54 @@ begin exact ⟨by simp only [of_real_re]; refl, by simp only [inner_self_nonneg_im, of_real_im]⟩ end -lemma inner_re_symm {x y : F} : re ⟪x, y⟫ = re ⟪y, x⟫ := +lemma inner_re_symm (x y : F) : re ⟪x, y⟫ = re ⟪y, x⟫ := by rw [←inner_conj_sym, conj_re] -lemma inner_im_symm {x y : F} : im ⟪x, y⟫ = -im ⟪y, x⟫ := +lemma inner_im_symm (x y : F) : im ⟪x, y⟫ = -im ⟪y, x⟫ := by rw [←inner_conj_sym, conj_im] -lemma inner_smul_left {x y : F} {r : 𝕜} : ⟪r • x, y⟫ = r† * ⟪x, y⟫ := +lemma inner_smul_left (x y : F) {r : 𝕜} : ⟪r • x, y⟫ = r† * ⟪x, y⟫ := c.smul_left _ _ _ -lemma inner_smul_right {x y : F} {r : 𝕜} : ⟪x, r • y⟫ = r * ⟪x, y⟫ := +lemma inner_smul_right (x y : F) {r : 𝕜} : ⟪x, r • y⟫ = r * ⟪x, y⟫ := by rw [←inner_conj_sym, inner_smul_left]; simp only [conj_conj, inner_conj_sym, ring_hom.map_mul] -lemma inner_zero_left {x : F} : ⟪0, x⟫ = 0 := +lemma inner_zero_left (x : F) : ⟪0, x⟫ = 0 := by rw [←zero_smul 𝕜 (0 : F), inner_smul_left]; simp only [zero_mul, ring_hom.map_zero] -lemma inner_zero_right {x : F} : ⟪x, 0⟫ = 0 := +lemma inner_zero_right (x : F) : ⟪x, 0⟫ = 0 := by rw [←inner_conj_sym, inner_zero_left]; simp only [ring_hom.map_zero] lemma inner_self_eq_zero {x : F} : ⟪x, x⟫ = 0 ↔ x = 0 := -iff.intro (c.definite _) (by { rintro rfl, exact inner_zero_left }) +iff.intro (c.definite _) (by { rintro rfl, exact inner_zero_left _ }) -lemma inner_self_re_to_K {x : F} : (re ⟪x, x⟫ : 𝕜) = ⟪x, x⟫ := +lemma inner_self_re_to_K (x : F) : (re ⟪x, x⟫ : 𝕜) = ⟪x, x⟫ := by norm_num [ext_iff, inner_self_nonneg_im] -lemma inner_abs_conj_sym {x y : F} : abs ⟪x, y⟫ = abs ⟪y, x⟫ := +lemma inner_abs_conj_sym (x y : F) : abs ⟪x, y⟫ = abs ⟪y, x⟫ := by rw [←inner_conj_sym, abs_conj] -lemma inner_neg_left {x y : F} : ⟪-x, y⟫ = -⟪x, y⟫ := +lemma inner_neg_left (x y : F) : ⟪-x, y⟫ = -⟪x, y⟫ := by { rw [← neg_one_smul 𝕜 x, inner_smul_left], simp } -lemma inner_neg_right {x y : F} : ⟪x, -y⟫ = -⟪x, y⟫ := +lemma inner_neg_right (x y : F) : ⟪x, -y⟫ = -⟪x, y⟫ := by rw [←inner_conj_sym, inner_neg_left]; simp only [ring_hom.map_neg, inner_conj_sym] -lemma inner_sub_left {x y z : F} : ⟪x - y, z⟫ = ⟪x, z⟫ - ⟪y, z⟫ := +lemma inner_sub_left (x y z : F) : ⟪x - y, z⟫ = ⟪x, z⟫ - ⟪y, z⟫ := by { simp [sub_eq_add_neg, inner_add_left, inner_neg_left] } -lemma inner_sub_right {x y z : F} : ⟪x, y - z⟫ = ⟪x, y⟫ - ⟪x, z⟫ := +lemma inner_sub_right (x y z : F) : ⟪x, y - z⟫ = ⟪x, y⟫ - ⟪x, z⟫ := by { simp [sub_eq_add_neg, inner_add_right, inner_neg_right] } -lemma inner_mul_conj_re_abs {x y : F} : re (⟪x, y⟫ * ⟪y, x⟫) = abs (⟪x, y⟫ * ⟪y, x⟫) := +lemma inner_mul_conj_re_abs (x y : F) : re (⟪x, y⟫ * ⟪y, x⟫) = abs (⟪x, y⟫ * ⟪y, x⟫) := by { rw [←inner_conj_sym, mul_comm], exact re_eq_abs_of_mul_conj (inner y x), } /-- Expand `inner (x + y) (x + y)` -/ -lemma inner_add_add_self {x y : F} : ⟪x + y, x + y⟫ = ⟪x, x⟫ + ⟪x, y⟫ + ⟪y, x⟫ + ⟪y, y⟫ := +lemma inner_add_add_self (x y : F) : ⟪x + y, x + y⟫ = ⟪x, x⟫ + ⟪x, y⟫ + ⟪y, x⟫ + ⟪y, y⟫ := by simp only [inner_add_left, inner_add_right]; ring /- Expand `inner (x - y) (x - y)` -/ -lemma inner_sub_sub_self {x y : F} : ⟪x - y, x - y⟫ = ⟪x, x⟫ - ⟪x, y⟫ - ⟪y, x⟫ + ⟪y, y⟫ := +lemma inner_sub_sub_self (x y : F) : ⟪x - y, x - y⟫ = ⟪x, x⟫ - ⟪x, y⟫ - ⟪y, x⟫ + ⟪y, y⟫ := by simp only [inner_sub_left, inner_sub_right]; ring /-- @@ -250,8 +250,8 @@ begin { change y ≠ 0 at hy, have hy' : ⟪y, y⟫ ≠ 0 := λ h, by rw [inner_self_eq_zero] at h; exact hy h, set T := ⟪y, x⟫ / ⟪y, y⟫ with hT, - have h₁ : re ⟪y, x⟫ = re ⟪x, y⟫ := inner_re_symm, - have h₂ : im ⟪y, x⟫ = -im ⟪x, y⟫ := inner_im_symm, + have h₁ : re ⟪y, x⟫ = re ⟪x, y⟫ := inner_re_symm _ _, + have h₂ : im ⟪y, x⟫ = -im ⟪x, y⟫ := inner_im_symm _ _, have h₃ : ⟪y, x⟫ * ⟪x, y⟫ * ⟪y, y⟫ / (⟪y, y⟫ * ⟪y, y⟫) = ⟪y, x⟫ * ⟪x, y⟫ / ⟪y, y⟫, { rw [mul_div_assoc], have : ⟪y, y⟫ / (⟪y, y⟫ * ⟪y, y⟫) = 1 / ⟪y, y⟫ := @@ -305,7 +305,7 @@ lemma inner_self_eq_norm_mul_norm (x : F) : re ⟪x, x⟫ = ‖x‖ * ‖x‖ := by rw [norm_eq_sqrt_inner, ←sqrt_mul inner_self_nonneg (re ⟪x, x⟫), sqrt_mul_self inner_self_nonneg] -lemma sqrt_norm_sq_eq_norm {x : F} : sqrt (norm_sqF x) = ‖x‖ := rfl +lemma sqrt_norm_sq_eq_norm (x : F) : sqrt (norm_sqF x) = ‖x‖ := rfl /-- Cauchy–Schwarz inequality with norm -/ lemma abs_inner_le_norm (x y : F) : abs ⟪x, y⟫ ≤ ‖x‖ * ‖y‖ := @@ -395,35 +395,36 @@ lemma real_inner_comm (x y : F) : ⟪y, x⟫_ℝ = ⟪x, y⟫_ℝ := @inner_conj lemma inner_eq_zero_sym {x y : E} : ⟪x, y⟫ = 0 ↔ ⟪y, x⟫ = 0 := ⟨λ h, by simp [←inner_conj_sym, h], λ h, by simp [←inner_conj_sym, h]⟩ -@[simp] lemma inner_self_nonneg_im {x : E} : im ⟪x, x⟫ = 0 := +@[simp] lemma inner_self_nonneg_im (x : E) : im ⟪x, x⟫ = 0 := by rw [← @of_real_inj 𝕜, im_eq_conj_sub]; simp -lemma inner_self_im_zero {x : E} : im ⟪x, x⟫ = 0 := inner_self_nonneg_im +lemma inner_self_im_zero (x : E) : im ⟪x, x⟫ = 0 := inner_self_nonneg_im _ -lemma inner_add_left {x y z : E} : ⟪x + y, z⟫ = ⟪x, z⟫ + ⟪y, z⟫ := +lemma inner_add_left (x y z : E) : ⟪x + y, z⟫ = ⟪x, z⟫ + ⟪y, z⟫ := inner_product_space.add_left _ _ _ -lemma inner_add_right {x y z : E} : ⟪x, y + z⟫ = ⟪x, y⟫ + ⟪x, z⟫ := +lemma inner_add_right (x y z : E) : ⟪x, y + z⟫ = ⟪x, y⟫ + ⟪x, z⟫ := by { rw [←inner_conj_sym, inner_add_left, ring_hom.map_add], simp only [inner_conj_sym] } -lemma inner_re_symm {x y : E} : re ⟪x, y⟫ = re ⟪y, x⟫ := +lemma inner_re_symm (x y : E) : re ⟪x, y⟫ = re ⟪y, x⟫ := by rw [←inner_conj_sym, conj_re] -lemma inner_im_symm {x y : E} : im ⟪x, y⟫ = -im ⟪y, x⟫ := +lemma inner_im_symm (x y : E) : im ⟪x, y⟫ = -im ⟪y, x⟫ := by rw [←inner_conj_sym, conj_im] -lemma inner_smul_left {x y : E} {r : 𝕜} : ⟪r • x, y⟫ = r† * ⟪x, y⟫ := +lemma inner_smul_left (x y : E) (r : 𝕜) : ⟪r • x, y⟫ = r† * ⟪x, y⟫ := inner_product_space.smul_left _ _ _ -lemma real_inner_smul_left {x y : F} {r : ℝ} : ⟪r • x, y⟫_ℝ = r * ⟪x, y⟫_ℝ := inner_smul_left +lemma real_inner_smul_left (x y : F) (r : ℝ) : ⟪r • x, y⟫_ℝ = r * ⟪x, y⟫_ℝ := inner_smul_left _ _ _ -lemma inner_smul_real_left {x y : E} {r : ℝ} : ⟪(r : 𝕜) • x, y⟫ = r • ⟪x, y⟫ := +lemma inner_smul_real_left (x y : E) (r : ℝ) : ⟪(r : 𝕜) • x, y⟫ = r • ⟪x, y⟫ := by { rw [inner_smul_left, conj_of_real, algebra.smul_def], refl } -lemma inner_smul_right {x y : E} {r : 𝕜} : ⟪x, r • y⟫ = r * ⟪x, y⟫ := +lemma inner_smul_right (x y : E) (r : 𝕜) : ⟪x, r • y⟫ = r * ⟪x, y⟫ := by rw [←inner_conj_sym, inner_smul_left, ring_hom.map_mul, conj_conj, inner_conj_sym] -lemma real_inner_smul_right {x y : F} {r : ℝ} : ⟪x, r • y⟫_ℝ = r * ⟪x, y⟫_ℝ := inner_smul_right +lemma real_inner_smul_right (x y : F) (r : ℝ) : ⟪x, r • y⟫_ℝ = r * ⟪x, y⟫_ℝ := +inner_smul_right _ _ _ -lemma inner_smul_real_right {x y : E} {r : ℝ} : ⟪x, (r : 𝕜) • y⟫ = r • ⟪x, y⟫ := +lemma inner_smul_real_right (x y : E) (r : ℝ) : ⟪x, (r : 𝕜) • y⟫ = r • ⟪x, y⟫ := by { rw [inner_smul_right, algebra.smul_def], refl } /-- The inner product as a sesquilinear form. @@ -433,19 +434,19 @@ Note that in the case `𝕜 = ℝ` this is a bilinear form. -/ def sesq_form_of_inner : E →ₗ[𝕜] E →ₗ⋆[𝕜] 𝕜 := linear_map.mk₂'ₛₗ (ring_hom.id 𝕜) (star_ring_end _) (λ x y, ⟪y, x⟫) - (λ x y z, inner_add_right) - (λ r x y, inner_smul_right) - (λ x y z, inner_add_left) - (λ r x y, inner_smul_left) + (λ x y z, inner_add_right _ _ _) + (λ r x y, inner_smul_right _ _ _) + (λ x y z, inner_add_left _ _ _) + (λ r x y, inner_smul_left _ _ _) /-- The real inner product as a bilinear form. -/ @[simps] def bilin_form_of_real_inner : bilin_form ℝ F := { bilin := inner, - bilin_add_left := λ x y z, inner_add_left, - bilin_smul_left := λ a x y, inner_smul_left, - bilin_add_right := λ x y z, inner_add_right, - bilin_smul_right := λ a x y, inner_smul_right } + bilin_add_left := inner_add_left, + bilin_smul_left := λ a x y, inner_smul_left _ _ _, + bilin_add_right := inner_add_right, + bilin_smul_right := λ a x y, inner_smul_right _ _ _ } /-- An inner product with a sum on the left. -/ lemma sum_inner {ι : Type*} (s : finset ι) (f : ι → E) (x : E) : @@ -480,16 +481,16 @@ lemma dfinsupp.inner_sum {ι : Type*} [dec : decidable_eq ι] {α : ι → Type* ⟪x, l.sum f⟫ = l.sum (λ i a, ⟪x, f i a⟫) := by simp only [dfinsupp.sum, inner_sum, smul_eq_mul] {contextual := tt} -@[simp] lemma inner_zero_left {x : E} : ⟪0, x⟫ = 0 := +@[simp] lemma inner_zero_left (x : E) : ⟪0, x⟫ = 0 := by rw [← zero_smul 𝕜 (0:E), inner_smul_left, ring_hom.map_zero, zero_mul] -lemma inner_re_zero_left {x : E} : re ⟪0, x⟫ = 0 := +lemma inner_re_zero_left (x : E) : re ⟪0, x⟫ = 0 := by simp only [inner_zero_left, add_monoid_hom.map_zero] -@[simp] lemma inner_zero_right {x : E} : ⟪x, 0⟫ = 0 := +@[simp] lemma inner_zero_right (x : E) : ⟪x, 0⟫ = 0 := by rw [←inner_conj_sym, inner_zero_left, ring_hom.map_zero] -lemma inner_re_zero_right {x : E} : re ⟪x, 0⟫ = 0 := +lemma inner_re_zero_right (x : E) : re ⟪x, 0⟫ = 0 := by simp only [inner_zero_right, add_monoid_hom.map_zero] lemma inner_self_nonneg {x : E} : 0 ≤ re ⟪x, x⟫ := @@ -506,7 +507,7 @@ begin rw [←norm_eq_zero], exact pow_eq_zero h₁ }, { rintro rfl, - exact inner_zero_left } + exact inner_zero_left _ } end @[simp] lemma inner_self_nonpos {x : E} : re ⟪x, x⟫ ≤ 0 ↔ x = 0 := @@ -525,7 +526,7 @@ end lemma real_inner_self_nonpos {x : F} : ⟪x, x⟫_ℝ ≤ 0 ↔ x = 0 := by { have h := @inner_self_nonpos ℝ F _ _ x, simpa using h } -@[simp] lemma inner_self_re_to_K {x : E} : (re ⟪x, x⟫ : 𝕜) = ⟪x, x⟫ := +@[simp] lemma inner_self_re_to_K (x : E) : (re ⟪x, x⟫ : 𝕜) = ⟪x, x⟫ := is_R_or_C.ext_iff.2 ⟨by simp only [of_real_re], by simp only [inner_self_nonneg_im, of_real_im]⟩ lemma inner_self_eq_norm_sq_to_K (x : E) : ⟪x, x⟫ = (‖x‖ ^ 2 : 𝕜) := @@ -535,48 +536,48 @@ begin exact_mod_cast (norm_sq_eq_inner x).symm end -lemma inner_self_re_abs {x : E} : re ⟪x, x⟫ = abs ⟪x, x⟫ := +lemma inner_self_re_abs (x : E) : re ⟪x, x⟫ = abs ⟪x, x⟫ := begin conv_rhs { rw [←inner_self_re_to_K] }, symmetry, exact is_R_or_C.abs_of_nonneg inner_self_nonneg, end -lemma inner_self_abs_to_K {x : E} : (absK ⟪x, x⟫ : 𝕜) = ⟪x, x⟫ := -by { rw [←inner_self_re_abs], exact inner_self_re_to_K } +lemma inner_self_abs_to_K (x : E) : (absK ⟪x, x⟫ : 𝕜) = ⟪x, x⟫ := +by { rw [←inner_self_re_abs], exact inner_self_re_to_K _ } -lemma real_inner_self_abs {x : F} : absR ⟪x, x⟫_ℝ = ⟪x, x⟫_ℝ := +lemma real_inner_self_abs (x : F) : absR ⟪x, x⟫_ℝ = ⟪x, x⟫_ℝ := by { have h := @inner_self_abs_to_K ℝ F _ _ x, simpa using h } -lemma inner_abs_conj_sym {x y : E} : abs ⟪x, y⟫ = abs ⟪y, x⟫ := +lemma inner_abs_conj_sym (x y : E) : abs ⟪x, y⟫ = abs ⟪y, x⟫ := by rw [←inner_conj_sym, abs_conj] -@[simp] lemma inner_neg_left {x y : E} : ⟪-x, y⟫ = -⟪x, y⟫ := +@[simp] lemma inner_neg_left (x y : E) : ⟪-x, y⟫ = -⟪x, y⟫ := by { rw [← neg_one_smul 𝕜 x, inner_smul_left], simp } -@[simp] lemma inner_neg_right {x y : E} : ⟪x, -y⟫ = -⟪x, y⟫ := +@[simp] lemma inner_neg_right (x y : E) : ⟪x, -y⟫ = -⟪x, y⟫ := by rw [←inner_conj_sym, inner_neg_left]; simp only [ring_hom.map_neg, inner_conj_sym] -lemma inner_neg_neg {x y : E} : ⟪-x, -y⟫ = ⟪x, y⟫ := by simp +lemma inner_neg_neg (x y : E) : ⟪-x, -y⟫ = ⟪x, y⟫ := by simp -@[simp] lemma inner_self_conj {x : E} : ⟪x, x⟫† = ⟪x, x⟫ := +@[simp] lemma inner_self_conj (x : E) : ⟪x, x⟫† = ⟪x, x⟫ := by rw [is_R_or_C.ext_iff]; exact ⟨by rw [conj_re], by rw [conj_im, inner_self_im_zero, neg_zero]⟩ -lemma inner_sub_left {x y z : E} : ⟪x - y, z⟫ = ⟪x, z⟫ - ⟪y, z⟫ := +lemma inner_sub_left (x y z : E) : ⟪x - y, z⟫ = ⟪x, z⟫ - ⟪y, z⟫ := by { simp [sub_eq_add_neg, inner_add_left] } -lemma inner_sub_right {x y z : E} : ⟪x, y - z⟫ = ⟪x, y⟫ - ⟪x, z⟫ := +lemma inner_sub_right (x y z : E) : ⟪x, y - z⟫ = ⟪x, y⟫ - ⟪x, z⟫ := by { simp [sub_eq_add_neg, inner_add_right] } -lemma inner_mul_conj_re_abs {x y : E} : re (⟪x, y⟫ * ⟪y, x⟫) = abs (⟪x, y⟫ * ⟪y, x⟫) := +lemma inner_mul_conj_re_abs (x y : E) : re (⟪x, y⟫ * ⟪y, x⟫) = abs (⟪x, y⟫ * ⟪y, x⟫) := by { rw [←inner_conj_sym, mul_comm], exact re_eq_abs_of_mul_conj (inner y x), } /-- Expand `⟪x + y, x + y⟫` -/ -lemma inner_add_add_self {x y : E} : ⟪x + y, x + y⟫ = ⟪x, x⟫ + ⟪x, y⟫ + ⟪y, x⟫ + ⟪y, y⟫ := +lemma inner_add_add_self (x y : E) : ⟪x + y, x + y⟫ = ⟪x, x⟫ + ⟪x, y⟫ + ⟪y, x⟫ + ⟪y, y⟫ := by simp only [inner_add_left, inner_add_right]; ring /-- Expand `⟪x + y, x + y⟫_ℝ` -/ -lemma real_inner_add_add_self {x y : F} : ⟪x + y, x + y⟫_ℝ = ⟪x, x⟫_ℝ + 2 * ⟪x, y⟫_ℝ + ⟪y, y⟫_ℝ := +lemma real_inner_add_add_self (x y : F) : ⟪x + y, x + y⟫_ℝ = ⟪x, x⟫_ℝ + 2 * ⟪x, y⟫_ℝ + ⟪y, y⟫_ℝ := begin have : ⟪y, x⟫_ℝ = ⟪x, y⟫_ℝ := by rw [←inner_conj_sym]; refl, simp only [inner_add_add_self, this, add_left_inj], @@ -584,11 +585,11 @@ begin end /- Expand `⟪x - y, x - y⟫` -/ -lemma inner_sub_sub_self {x y : E} : ⟪x - y, x - y⟫ = ⟪x, x⟫ - ⟪x, y⟫ - ⟪y, x⟫ + ⟪y, y⟫ := +lemma inner_sub_sub_self (x y : E) : ⟪x - y, x - y⟫ = ⟪x, x⟫ - ⟪x, y⟫ - ⟪y, x⟫ + ⟪y, y⟫ := by simp only [inner_sub_left, inner_sub_right]; ring /-- Expand `⟪x - y, x - y⟫_ℝ` -/ -lemma real_inner_sub_sub_self {x y : F} : ⟪x - y, x - y⟫_ℝ = ⟪x, x⟫_ℝ - 2 * ⟪x, y⟫_ℝ + ⟪y, y⟫_ℝ := +lemma real_inner_sub_sub_self (x y : F) : ⟪x - y, x - y⟫_ℝ = ⟪x, x⟫_ℝ - 2 * ⟪x, y⟫_ℝ + ⟪y, y⟫_ℝ := begin have : ⟪y, x⟫_ℝ = ⟪x, y⟫_ℝ := by rw [←inner_conj_sym]; refl, simp only [inner_sub_sub_self, this, add_left_inj], @@ -619,14 +620,14 @@ begin { rw [hy], simp only [is_R_or_C.abs_zero, inner_zero_left, mul_zero, add_monoid_hom.map_zero] }, { have hy' : ⟪y, y⟫ ≠ 0 := inner_self_eq_zero.not.2 hy, set T := ⟪y, x⟫ / ⟪y, y⟫ with hT, - have h₁ : re ⟪y, x⟫ = re ⟪x, y⟫ := inner_re_symm, - have h₂ : im ⟪y, x⟫ = -im ⟪x, y⟫ := inner_im_symm, + have h₁ : re ⟪y, x⟫ = re ⟪x, y⟫ := inner_re_symm _ _, + have h₂ : im ⟪y, x⟫ = -im ⟪x, y⟫ := inner_im_symm _ _, have h₃ : ⟪y, x⟫ * ⟪x, y⟫ * ⟪y, y⟫ / (⟪y, y⟫ * ⟪y, y⟫) = ⟪y, x⟫ * ⟪x, y⟫ / ⟪y, y⟫, { rw [mul_div_assoc], have : ⟪y, y⟫ / (⟪y, y⟫ * ⟪y, y⟫) = 1 / ⟪y, y⟫ := by rw [div_mul_eq_div_mul_one_div, div_self hy', one_mul], rw [this, div_eq_mul_inv, one_mul, ←div_eq_mul_inv] }, - have h₄ : ⟪y, y⟫ = re ⟪y, y⟫ := inner_self_re_to_K.symm, + have h₄ : ⟪y, y⟫ = re ⟪y, y⟫ := (inner_self_re_to_K _).symm, have h₅ : re ⟪y, y⟫ > 0, { refine lt_of_le_of_ne inner_self_nonneg _, intro H, @@ -973,7 +974,7 @@ lemma real_inner_self_eq_norm_sq (x : F) : ⟪x, x⟫_ℝ = ‖x‖^2 := by rw [pow_two, real_inner_self_eq_norm_mul_norm] /-- Expand the square -/ -lemma norm_add_sq {x y : E} : ‖x + y‖^2 = ‖x‖^2 + 2 * (re ⟪x, y⟫) + ‖y‖^2 := +lemma norm_add_sq (x y : E) : ‖x + y‖^2 = ‖x‖^2 + 2 * (re ⟪x, y⟫) + ‖y‖^2 := begin repeat {rw [sq, ←inner_self_eq_norm_mul_norm]}, rw [inner_add_add_self, two_mul], @@ -984,21 +985,21 @@ end alias norm_add_sq ← norm_add_pow_two /-- Expand the square -/ -lemma norm_add_sq_real {x y : F} : ‖x + y‖^2 = ‖x‖^2 + 2 * ⟪x, y⟫_ℝ + ‖y‖^2 := -by { have h := @norm_add_sq ℝ F _ _, simpa using h } +lemma norm_add_sq_real (x y : F) : ‖x + y‖^2 = ‖x‖^2 + 2 * ⟪x, y⟫_ℝ + ‖y‖^2 := +by { have h := @norm_add_sq ℝ _ _ _ x y, simpa using h } alias norm_add_sq_real ← norm_add_pow_two_real /-- Expand the square -/ -lemma norm_add_mul_self {x y : E} : ‖x + y‖ * ‖x + y‖ = ‖x‖ * ‖x‖ + 2 * (re ⟪x, y⟫) + ‖y‖ * ‖y‖ := -by { repeat {rw [← sq]}, exact norm_add_sq } +lemma norm_add_mul_self (x y : E) : ‖x + y‖ * ‖x + y‖ = ‖x‖ * ‖x‖ + 2 * (re ⟪x, y⟫) + ‖y‖ * ‖y‖ := +by { repeat {rw [← sq]}, exact norm_add_sq _ _ } /-- Expand the square -/ -lemma norm_add_mul_self_real {x y : F} : ‖x + y‖ * ‖x + y‖ = ‖x‖ * ‖x‖ + 2 * ⟪x, y⟫_ℝ + ‖y‖ * ‖y‖ := -by { have h := @norm_add_mul_self ℝ F _ _, simpa using h } +lemma norm_add_mul_self_real (x y : F) : ‖x + y‖ * ‖x + y‖ = ‖x‖ * ‖x‖ + 2 * ⟪x, y⟫_ℝ + ‖y‖ * ‖y‖ := +by { have h := @norm_add_mul_self ℝ _ _ _ x y, simpa using h } /-- Expand the square -/ -lemma norm_sub_sq {x y : E} : ‖x - y‖^2 = ‖x‖^2 - 2 * (re ⟪x, y⟫) + ‖y‖^2 := +lemma norm_sub_sq (x y : E) : ‖x - y‖^2 = ‖x‖^2 - 2 * (re ⟪x, y⟫) + ‖y‖^2 := begin repeat {rw [sq, ←inner_self_eq_norm_mul_norm]}, rw [inner_sub_sub_self], @@ -1014,18 +1015,18 @@ end alias norm_sub_sq ← norm_sub_pow_two /-- Expand the square -/ -lemma norm_sub_sq_real {x y : F} : ‖x - y‖^2 = ‖x‖^2 - 2 * ⟪x, y⟫_ℝ + ‖y‖^2 := -norm_sub_sq +lemma norm_sub_sq_real (x y : F) : ‖x - y‖^2 = ‖x‖^2 - 2 * ⟪x, y⟫_ℝ + ‖y‖^2 := +norm_sub_sq _ _ alias norm_sub_sq_real ← norm_sub_pow_two_real /-- Expand the square -/ -lemma norm_sub_mul_self {x y : E} : ‖x - y‖ * ‖x - y‖ = ‖x‖ * ‖x‖ - 2 * re ⟪x, y⟫ + ‖y‖ * ‖y‖ := -by { repeat {rw [← sq]}, exact norm_sub_sq } +lemma norm_sub_mul_self (x y : E) : ‖x - y‖ * ‖x - y‖ = ‖x‖ * ‖x‖ - 2 * re ⟪x, y⟫ + ‖y‖ * ‖y‖ := +by { repeat {rw [← sq]}, exact norm_sub_sq _ _ } /-- Expand the square -/ -lemma norm_sub_mul_self_real {x y : F} : ‖x - y‖ * ‖x - y‖ = ‖x‖ * ‖x‖ - 2 * ⟪x, y⟫_ℝ + ‖y‖ * ‖y‖ := -by { have h := @norm_sub_mul_self ℝ F _ _, simpa using h } +lemma norm_sub_mul_self_real (x y : F) : ‖x - y‖ * ‖x - y‖ = ‖x‖ * ‖x‖ - 2 * ⟪x, y⟫_ℝ + ‖y‖ * ‖y‖ := +by { have h := @norm_sub_mul_self ℝ _ _ _ x y, simpa using h } /-- Cauchy–Schwarz inequality with norm -/ lemma abs_inner_le_norm (x y : E) : abs ⟪x, y⟫ ≤ ‖x‖ * ‖y‖ := @@ -1690,8 +1691,8 @@ by simp_rw [sum_inner, inner_sum, real_inner_smul_left, real_inner_smul_right, /-- The inner product as a sesquilinear map. -/ def innerₛₗ : E →ₗ⋆[𝕜] E →ₗ[𝕜] 𝕜 := -linear_map.mk₂'ₛₗ _ _ (λ v w, ⟪v, w⟫) (λ _ _ _, inner_add_left) (λ _ _ _, inner_smul_left) -(λ _ _ _, inner_add_right) (λ _ _ _, inner_smul_right) +linear_map.mk₂'ₛₗ _ _ (λ v w, ⟪v, w⟫) inner_add_left (λ _ _ _, inner_smul_left _ _ _) + inner_add_right (λ _ _ _, inner_smul_right _ _ _) @[simp] lemma innerₛₗ_apply_coe (v : E) : (innerₛₗ v : E → 𝕜) = λ w, ⟪v, w⟫ := rfl @@ -1730,7 +1731,7 @@ def innerSL_flip : E →L[𝕜] E →L⋆[𝕜] 𝕜 := @continuous_linear_map.flipₗᵢ' 𝕜 𝕜 𝕜 E E 𝕜 _ _ _ _ _ _ _ _ _ (ring_hom.id 𝕜) (star_ring_end 𝕜) _ _ innerSL -@[simp] lemma innerSL_flip_apply {x y : E} : innerSL_flip x y = ⟪y, x⟫ := rfl +@[simp] lemma innerSL_flip_apply (x y : E) : innerSL_flip x y = ⟪y, x⟫ := rfl namespace continuous_linear_map @@ -1769,10 +1770,10 @@ instance may be not definitionally equal to some other “natural” instance. S -/ lemma is_bounded_bilinear_map_inner [normed_space ℝ E] : is_bounded_bilinear_map ℝ (λ p : E × E, ⟪p.1, p.2⟫) := -{ add_left := λ _ _ _, inner_add_left, +{ add_left := inner_add_left, smul_left := λ r x y, by simp only [← algebra_map_smul 𝕜 r x, algebra_map_eq_of_real, inner_smul_real_left], - add_right := λ _ _ _, inner_add_right, + add_right := inner_add_right, smul_right := λ r x y, by simp only [← algebra_map_smul 𝕜 r y, algebra_map_eq_of_real, inner_smul_real_right], bound := ⟨1, zero_lt_one, λ x y, @@ -1848,8 +1849,8 @@ instance submodule.inner_product_space (W : submodule 𝕜 E) : inner_product_sp inner := λ x y, ⟪(x:E), (y:E)⟫, conj_sym := λ _ _, inner_conj_sym _ _ , norm_sq_eq_inner := λ _, norm_sq_eq_inner _, - add_left := λ _ _ _ , inner_add_left, - smul_left := λ _ _ _, inner_smul_left, + add_left := λ _ _ _, inner_add_left _ _ _, + smul_left := λ _ _ _, inner_smul_left _ _ _, ..submodule.normed_space W } /-- The inner product on submodules is the same as on the ambient space. -/ @@ -2102,7 +2103,7 @@ structure. -/ def inner_product_space.is_R_or_C_to_real : inner_product_space ℝ E := { to_normed_add_comm_group := inner_product_space.to_normed_add_comm_group 𝕜, norm_sq_eq_inner := norm_sq_eq_inner, - conj_sym := λ x y, inner_re_symm, + conj_sym := λ x y, inner_re_symm _ _, add_left := λ x y z, by { change re ⟪x + y, z⟫ = re ⟪x, z⟫ + re ⟪y, z⟫, simp only [inner_add_left, map_add] }, @@ -2208,7 +2209,7 @@ variables (K : submodule 𝕜 E) /-- The subspace of vectors orthogonal to a given subspace. -/ def submodule.orthogonal : submodule 𝕜 E := { carrier := {v | ∀ u ∈ K, ⟪u, v⟫ = 0}, - zero_mem' := λ _ _, inner_zero_right, + zero_mem' := λ _ _, inner_zero_right _, add_mem' := λ x y hx hy u hu, by rw [inner_add_right, hx u hu, hy u hu, add_zero], smul_mem' := λ c x hx u hu, by rw [inner_smul_right, hx u hu, mul_zero] } @@ -2392,8 +2393,8 @@ protected lemma continuous_inner : begin let inner' : E →+ E →+ 𝕜 := { to_fun := λ x, (innerₛₗ x).to_add_monoid_hom, - map_zero' := by ext x; exact inner_zero_left, - map_add' := λ x y, by ext z; exact inner_add_left }, + map_zero' := by ext x; exact inner_zero_left _, + map_add' := λ x y, by ext z; exact inner_add_left _ _ _ }, have : continuous (λ p : E × E, inner' p.1 p.2) := continuous_inner, rw [completion.has_inner, uncurry_curry _], change continuous (((dense_inducing_to_compl E).prod (dense_inducing_to_compl E)).extend diff --git a/src/analysis/inner_product_space/lax_milgram.lean b/src/analysis/inner_product_space/lax_milgram.lean index b7ed91982ec9e..24753555f05ef 100644 --- a/src/analysis/inner_product_space/lax_milgram.lean +++ b/src/analysis/inner_product_space/lax_milgram.lean @@ -101,7 +101,7 @@ begin ... = ⟪B♯ w, w⟫_ℝ : (continuous_linear_map_of_bilin_apply ℝ B w w).symm ... = 0 : mem_w_orthogonal _ ⟨w, rfl⟩ }, { exact mul_nonneg (mul_nonneg C_pos.le (norm_nonneg w)) (norm_nonneg w) } }, - exact inner_zero_left, + exact inner_zero_left _, end /-- diff --git a/src/analysis/inner_product_space/projection.lean b/src/analysis/inner_product_space/projection.lean index 916245864d944..49b4476a27ee1 100644 --- a/src/analysis/inner_product_space/projection.lean +++ b/src/analysis/inner_product_space/projection.lean @@ -246,7 +246,7 @@ begin ‖u - v‖ * ‖u - v‖ ≤ ‖u - v‖ * ‖u - v‖ - 2 * inner (u - v) ((w:F) - v) : by linarith ... ≤ ‖u - v‖^2 - 2 * inner (u - v) ((w:F) - v) + ‖(w:F) - v‖^2 : by { rw sq, refine le_add_of_nonneg_right _, exact sq_nonneg _ } - ... = ‖(u - v) - (w - v)‖^2 : norm_sub_sq.symm + ... = ‖(u - v) - (w - v)‖^2 : (norm_sub_sq _ _).symm ... = ‖u - w‖ * ‖u - w‖ : by { have : (u - v) - (w - v) = u - w, abel, rw [this, sq] } }, { show (⨅ (w : K), ‖u - w‖) ≤ (λw:K, ‖u - w‖) ⟨v, hv⟩, diff --git a/src/analysis/inner_product_space/symmetric.lean b/src/analysis/inner_product_space/symmetric.lean index 236b19b900c36..aa19ff3829314 100644 --- a/src/analysis/inner_product_space/symmetric.lean +++ b/src/analysis/inner_product_space/symmetric.lean @@ -69,7 +69,7 @@ by rw [hT x y, inner_conj_sym] hT x y lemma is_symmetric_zero : (0 : E →ₗ[𝕜] E).is_symmetric := -λ x y, (inner_zero_right : ⟪x, 0⟫ = 0).symm ▸ (inner_zero_left : ⟪0, y⟫ = 0) +λ x y, (inner_zero_right x : ⟪x, 0⟫ = 0).symm ▸ (inner_zero_left y : ⟪0, y⟫ = 0) lemma is_symmetric_id : (linear_map.id : E →ₗ[𝕜] E).is_symmetric := λ x y, rfl @@ -94,7 +94,7 @@ begin by { intro k, rw [←T.map_sub, hT] }, refine tendsto_nhds_unique ((hTu.sub_const _).inner tendsto_const_nhds) _, simp_rw hlhs, - rw ←@inner_zero_left 𝕜 E _ _ (T (y - T x)), + rw ←inner_zero_left (T (y - T x)), refine filter.tendsto.inner _ tendsto_const_nhds, rw ←sub_self x, exact hu.sub_const _, diff --git a/src/measure_theory/function/l2_space.lean b/src/measure_theory/function/l2_space.lean index 1af86e1ce6aa7..3f9ecafe7dd00 100644 --- a/src/measure_theory/function/l2_space.lean +++ b/src/measure_theory/function/l2_space.lean @@ -195,7 +195,7 @@ begin from indicator_const_Lp_coe_fn_nmem, refine h_indicator.mono (λ x hx hxs, _), rw hx hxs, - exact inner_zero_left, }, + exact inner_zero_left _, }, rw [h_left, h_right, add_zero], end From 3fc0b254310908f70a1a75f01147d52e53e9f8a2 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 14 Mar 2023 15:36:40 +0000 Subject: [PATCH 0625/1291] chore(analysis/inner_product_space): rename `_sym` to `_symm` (#18580) This is the preferred spelling of `symmetric` in mathlib. `sym` refers to the type of unordered pairs. --- src/analysis/inner_product_space/adjoint.lean | 2 +- src/analysis/inner_product_space/basic.lean | 96 +++++++++---------- src/analysis/inner_product_space/dual.lean | 8 +- .../gram_schmidt_ortho.lean | 4 +- .../inner_product_space/l2_space.lean | 6 +- src/analysis/inner_product_space/pi_L2.lean | 4 +- .../inner_product_space/projection.lean | 14 +-- .../inner_product_space/symmetric.lean | 4 +- src/analysis/quaternion.lean | 2 +- src/linear_algebra/matrix/ldl.lean | 2 +- src/linear_algebra/matrix/pos_def.lean | 2 +- src/measure_theory/function/l2_space.lean | 2 +- 12 files changed, 73 insertions(+), 73 deletions(-) diff --git a/src/analysis/inner_product_space/adjoint.lean b/src/analysis/inner_product_space/adjoint.lean index 044b0de10850d..f5a26f01b4cbc 100644 --- a/src/analysis/inner_product_space/adjoint.lean +++ b/src/analysis/inner_product_space/adjoint.lean @@ -70,7 +70,7 @@ by { simp only [adjoint_aux_apply, to_dual_symm_apply, to_sesq_form_apply_coe, c innerSL_apply_coe]} lemma adjoint_aux_inner_right (A : E →L[𝕜] F) (x : E) (y : F) : ⟪x, adjoint_aux A y⟫ = ⟪A x, y⟫ := -by rw [←inner_conj_sym, adjoint_aux_inner_left, inner_conj_sym] +by rw [←inner_conj_symm, adjoint_aux_inner_left, inner_conj_symm] variables [complete_space F] diff --git a/src/analysis/inner_product_space/basic.lean b/src/analysis/inner_product_space/basic.lean index 0684b1c478e6b..818d75a5cda04 100644 --- a/src/analysis/inner_product_space/basic.lean +++ b/src/analysis/inner_product_space/basic.lean @@ -102,7 +102,7 @@ To construct a norm from an inner product, see `inner_product_space.of_core`. class inner_product_space (𝕜 : Type*) (E : Type*) [is_R_or_C 𝕜] extends normed_add_comm_group E, normed_space 𝕜 E, has_inner 𝕜 E := (norm_sq_eq_inner : ∀ (x : E), ‖x‖^2 = re (inner x x)) -(conj_sym : ∀ x y, conj (inner y x) = inner x y) +(conj_symm : ∀ x y, conj (inner y x) = inner x y) (add_left : ∀ x y z, inner (x + y) z = inner x z + inner y z) (smul_left : ∀ x y r, inner (r • x) y = (conj r) * inner x y) @@ -134,7 +134,7 @@ structure inner_product_space.core (𝕜 : Type*) (F : Type*) [is_R_or_C 𝕜] [add_comm_group F] [module 𝕜 F] := (inner : F → F → 𝕜) -(conj_sym : ∀ x y, conj (inner y x) = inner x y) +(conj_symm : ∀ x y, conj (inner y x) = inner x y) (nonneg_re : ∀ x, 0 ≤ re (inner x x)) (definite : ∀ x, inner x x = 0 → x = 0) (add_left : ∀ x y z, inner (x + y) z = inner x z + inner y z) @@ -166,12 +166,12 @@ def norm_sq (x : F) := reK ⟪x, x⟫ local notation `norm_sqF` := @norm_sq 𝕜 F _ _ _ _ -lemma inner_conj_sym (x y : F) : ⟪y, x⟫† = ⟪x, y⟫ := c.conj_sym x y +lemma inner_conj_symm (x y : F) : ⟪y, x⟫† = ⟪x, y⟫ := c.conj_symm x y lemma inner_self_nonneg {x : F} : 0 ≤ re ⟪x, x⟫ := c.nonneg_re _ lemma inner_self_nonneg_im (x : F) : im ⟪x, x⟫ = 0 := -by rw [← @of_real_inj 𝕜, im_eq_conj_sub]; simp [inner_conj_sym] +by rw [← @of_real_inj 𝕜, im_eq_conj_sub]; simp [inner_conj_symm] lemma inner_self_im_zero (x : F) : im ⟪x, x⟫ = 0 := inner_self_nonneg_im _ @@ -180,7 +180,7 @@ lemma inner_add_left (x y z : F) : ⟪x + y, z⟫ = ⟪x, z⟫ + ⟪y, z⟫ := c.add_left _ _ _ lemma inner_add_right (x y z : F) : ⟪x, y + z⟫ = ⟪x, y⟫ + ⟪x, z⟫ := -by rw [←inner_conj_sym, inner_add_left, ring_hom.map_add]; simp only [inner_conj_sym] +by rw [←inner_conj_symm, inner_add_left, ring_hom.map_add]; simp only [inner_conj_symm] lemma inner_norm_sq_eq_inner_self (x : F) : (norm_sqF x : 𝕜) = ⟪x, x⟫ := begin @@ -189,22 +189,22 @@ begin end lemma inner_re_symm (x y : F) : re ⟪x, y⟫ = re ⟪y, x⟫ := -by rw [←inner_conj_sym, conj_re] +by rw [←inner_conj_symm, conj_re] lemma inner_im_symm (x y : F) : im ⟪x, y⟫ = -im ⟪y, x⟫ := -by rw [←inner_conj_sym, conj_im] +by rw [←inner_conj_symm, conj_im] lemma inner_smul_left (x y : F) {r : 𝕜} : ⟪r • x, y⟫ = r† * ⟪x, y⟫ := c.smul_left _ _ _ lemma inner_smul_right (x y : F) {r : 𝕜} : ⟪x, r • y⟫ = r * ⟪x, y⟫ := -by rw [←inner_conj_sym, inner_smul_left]; simp only [conj_conj, inner_conj_sym, ring_hom.map_mul] +by rw [←inner_conj_symm, inner_smul_left]; simp only [conj_conj, inner_conj_symm, ring_hom.map_mul] lemma inner_zero_left (x : F) : ⟪0, x⟫ = 0 := by rw [←zero_smul 𝕜 (0 : F), inner_smul_left]; simp only [zero_mul, ring_hom.map_zero] lemma inner_zero_right (x : F) : ⟪x, 0⟫ = 0 := -by rw [←inner_conj_sym, inner_zero_left]; simp only [ring_hom.map_zero] +by rw [←inner_conj_symm, inner_zero_left]; simp only [ring_hom.map_zero] lemma inner_self_eq_zero {x : F} : ⟪x, x⟫ = 0 ↔ x = 0 := iff.intro (c.definite _) (by { rintro rfl, exact inner_zero_left _ }) @@ -212,14 +212,14 @@ iff.intro (c.definite _) (by { rintro rfl, exact inner_zero_left _ }) lemma inner_self_re_to_K (x : F) : (re ⟪x, x⟫ : 𝕜) = ⟪x, x⟫ := by norm_num [ext_iff, inner_self_nonneg_im] -lemma inner_abs_conj_sym (x y : F) : abs ⟪x, y⟫ = abs ⟪y, x⟫ := -by rw [←inner_conj_sym, abs_conj] +lemma inner_abs_conj_symm (x y : F) : abs ⟪x, y⟫ = abs ⟪y, x⟫ := +by rw [←inner_conj_symm, abs_conj] lemma inner_neg_left (x y : F) : ⟪-x, y⟫ = -⟪x, y⟫ := by { rw [← neg_one_smul 𝕜 x, inner_smul_left], simp } lemma inner_neg_right (x y : F) : ⟪x, -y⟫ = -⟪x, y⟫ := -by rw [←inner_conj_sym, inner_neg_left]; simp only [ring_hom.map_neg, inner_conj_sym] +by rw [←inner_conj_symm, inner_neg_left]; simp only [ring_hom.map_neg, inner_conj_symm] lemma inner_sub_left (x y z : F) : ⟪x - y, z⟫ = ⟪x, z⟫ - ⟪y, z⟫ := by { simp [sub_eq_add_neg, inner_add_left, inner_neg_left] } @@ -228,7 +228,7 @@ lemma inner_sub_right (x y z : F) : ⟪x, y - z⟫ = ⟪x, y⟫ - ⟪x, z⟫ := by { simp [sub_eq_add_neg, inner_add_right, inner_neg_right] } lemma inner_mul_conj_re_abs (x y : F) : re (⟪x, y⟫ * ⟪y, x⟫) = abs (⟪x, y⟫ * ⟪y, x⟫) := -by { rw [←inner_conj_sym, mul_comm], exact re_eq_abs_of_mul_conj (inner y x), } +by { rw [←inner_conj_symm, mul_comm], exact re_eq_abs_of_mul_conj (inner y x), } /-- Expand `inner (x + y) (x + y)` -/ lemma inner_add_add_self (x y : F) : ⟪x + y, x + y⟫ = ⟪x, x⟫ + ⟪x, y⟫ + ⟪y, x⟫ + ⟪y, y⟫ := @@ -276,7 +276,7 @@ begin ... = re ⟪x, x⟫ - re (T† * ⟪y, x⟫) - re (T * ⟪x, y⟫) + re (T * T† * ⟪y, y⟫) : by simp only [inner_smul_left, inner_smul_right, mul_assoc] ... = re ⟪x, x⟫ - re (⟪x, y⟫ / ⟪y, y⟫ * ⟪y, x⟫) - : by field_simp [-mul_re, inner_conj_sym, hT, map_div₀, h₁, h₃] + : by field_simp [-mul_re, inner_conj_symm, hT, map_div₀, h₁, h₃] ... = re ⟪x, x⟫ - re (⟪x, y⟫ * ⟪y, x⟫ / ⟪y, y⟫) : by rw ←mul_div_right_comm ... = re ⟪x, x⟫ - re (⟪x, y⟫ * ⟪y, x⟫ / re ⟪y, y⟫) @@ -316,7 +316,7 @@ begin rw H, conv begin - to_lhs, congr, rw [inner_abs_conj_sym], + to_lhs, congr, rw [inner_abs_conj_symm], end, exact inner_mul_inner_self_le y x, end @@ -331,7 +331,7 @@ add_group_norm.to_normed_add_comm_group have h₁ : abs ⟪x, y⟫ ≤ ‖x‖ * ‖y‖ := abs_inner_le_norm _ _, have h₂ : re ⟪x, y⟫ ≤ abs ⟪x, y⟫ := re_le_abs _, have h₃ : re ⟪x, y⟫ ≤ ‖x‖ * ‖y‖ := by linarith, - have h₄ : re ⟪y, x⟫ ≤ ‖x‖ * ‖y‖ := by rwa [←inner_conj_sym, conj_re], + have h₄ : re ⟪y, x⟫ ≤ ‖x‖ * ‖y‖ := by rwa [←inner_conj_symm, conj_re], have : ‖x + y‖ * ‖x + y‖ ≤ (‖x‖ + ‖y‖) * (‖x‖ + ‖y‖), { simp only [←inner_self_eq_norm_mul_norm, inner_add_add_self, mul_add, mul_comm, map_add], linarith }, @@ -389,11 +389,11 @@ export inner_product_space (norm_sq_eq_inner) section basic_properties -@[simp] lemma inner_conj_sym (x y : E) : ⟪y, x⟫† = ⟪x, y⟫ := inner_product_space.conj_sym _ _ -lemma real_inner_comm (x y : F) : ⟪y, x⟫_ℝ = ⟪x, y⟫_ℝ := @inner_conj_sym ℝ _ _ _ x y +@[simp] lemma inner_conj_symm (x y : E) : ⟪y, x⟫† = ⟪x, y⟫ := inner_product_space.conj_symm _ _ +lemma real_inner_comm (x y : F) : ⟪y, x⟫_ℝ = ⟪x, y⟫_ℝ := @inner_conj_symm ℝ _ _ _ x y -lemma inner_eq_zero_sym {x y : E} : ⟪x, y⟫ = 0 ↔ ⟪y, x⟫ = 0 := -⟨λ h, by simp [←inner_conj_sym, h], λ h, by simp [←inner_conj_sym, h]⟩ +lemma inner_eq_zero_symm {x y : E} : ⟪x, y⟫ = 0 ↔ ⟪y, x⟫ = 0 := +⟨λ h, by simp [←inner_conj_symm, h], λ h, by simp [←inner_conj_symm, h]⟩ @[simp] lemma inner_self_nonneg_im (x : E) : im ⟪x, x⟫ = 0 := by rw [← @of_real_inj 𝕜, im_eq_conj_sub]; simp @@ -404,13 +404,13 @@ lemma inner_add_left (x y z : E) : ⟪x + y, z⟫ = ⟪x, z⟫ + ⟪y, z⟫ := inner_product_space.add_left _ _ _ lemma inner_add_right (x y z : E) : ⟪x, y + z⟫ = ⟪x, y⟫ + ⟪x, z⟫ := -by { rw [←inner_conj_sym, inner_add_left, ring_hom.map_add], simp only [inner_conj_sym] } +by { rw [←inner_conj_symm, inner_add_left, ring_hom.map_add], simp only [inner_conj_symm] } lemma inner_re_symm (x y : E) : re ⟪x, y⟫ = re ⟪y, x⟫ := -by rw [←inner_conj_sym, conj_re] +by rw [←inner_conj_symm, conj_re] lemma inner_im_symm (x y : E) : im ⟪x, y⟫ = -im ⟪y, x⟫ := -by rw [←inner_conj_sym, conj_im] +by rw [←inner_conj_symm, conj_im] lemma inner_smul_left (x y : E) (r : 𝕜) : ⟪r • x, y⟫ = r† * ⟪x, y⟫ := inner_product_space.smul_left _ _ _ @@ -420,7 +420,7 @@ lemma inner_smul_real_left (x y : E) (r : ℝ) : ⟪(r : 𝕜) • x, y⟫ = r by { rw [inner_smul_left, conj_of_real, algebra.smul_def], refl } lemma inner_smul_right (x y : E) (r : 𝕜) : ⟪x, r • y⟫ = r * ⟪x, y⟫ := -by rw [←inner_conj_sym, inner_smul_left, ring_hom.map_mul, conj_conj, inner_conj_sym] +by rw [←inner_conj_symm, inner_smul_left, ring_hom.map_mul, conj_conj, inner_conj_symm] lemma real_inner_smul_right (x y : F) (r : ℝ) : ⟪x, r • y⟫_ℝ = r * ⟪x, y⟫_ℝ := inner_smul_right _ _ _ @@ -488,7 +488,7 @@ lemma inner_re_zero_left (x : E) : re ⟪0, x⟫ = 0 := by simp only [inner_zero_left, add_monoid_hom.map_zero] @[simp] lemma inner_zero_right (x : E) : ⟪x, 0⟫ = 0 := -by rw [←inner_conj_sym, inner_zero_left, ring_hom.map_zero] +by rw [←inner_conj_symm, inner_zero_left, ring_hom.map_zero] lemma inner_re_zero_right (x : E) : re ⟪x, 0⟫ = 0 := by simp only [inner_zero_right, add_monoid_hom.map_zero] @@ -549,14 +549,14 @@ by { rw [←inner_self_re_abs], exact inner_self_re_to_K _ } lemma real_inner_self_abs (x : F) : absR ⟪x, x⟫_ℝ = ⟪x, x⟫_ℝ := by { have h := @inner_self_abs_to_K ℝ F _ _ x, simpa using h } -lemma inner_abs_conj_sym (x y : E) : abs ⟪x, y⟫ = abs ⟪y, x⟫ := -by rw [←inner_conj_sym, abs_conj] +lemma inner_abs_conj_symm (x y : E) : abs ⟪x, y⟫ = abs ⟪y, x⟫ := +by rw [←inner_conj_symm, abs_conj] @[simp] lemma inner_neg_left (x y : E) : ⟪-x, y⟫ = -⟪x, y⟫ := by { rw [← neg_one_smul 𝕜 x, inner_smul_left], simp } @[simp] lemma inner_neg_right (x y : E) : ⟪x, -y⟫ = -⟪x, y⟫ := -by rw [←inner_conj_sym, inner_neg_left]; simp only [ring_hom.map_neg, inner_conj_sym] +by rw [←inner_conj_symm, inner_neg_left]; simp only [ring_hom.map_neg, inner_conj_symm] lemma inner_neg_neg (x y : E) : ⟪-x, -y⟫ = ⟪x, y⟫ := by simp @@ -570,7 +570,7 @@ lemma inner_sub_right (x y z : E) : ⟪x, y - z⟫ = ⟪x, y⟫ - ⟪x, z⟫ := by { simp [sub_eq_add_neg, inner_add_right] } lemma inner_mul_conj_re_abs (x y : E) : re (⟪x, y⟫ * ⟪y, x⟫) = abs (⟪x, y⟫ * ⟪y, x⟫) := -by { rw [←inner_conj_sym, mul_comm], exact re_eq_abs_of_mul_conj (inner y x), } +by { rw [←inner_conj_symm, mul_comm], exact re_eq_abs_of_mul_conj (inner y x), } /-- Expand `⟪x + y, x + y⟫` -/ lemma inner_add_add_self (x y : E) : ⟪x + y, x + y⟫ = ⟪x, x⟫ + ⟪x, y⟫ + ⟪y, x⟫ + ⟪y, y⟫ := @@ -579,7 +579,7 @@ by simp only [inner_add_left, inner_add_right]; ring /-- Expand `⟪x + y, x + y⟫_ℝ` -/ lemma real_inner_add_add_self (x y : F) : ⟪x + y, x + y⟫_ℝ = ⟪x, x⟫_ℝ + 2 * ⟪x, y⟫_ℝ + ⟪y, y⟫_ℝ := begin - have : ⟪y, x⟫_ℝ = ⟪x, y⟫_ℝ := by rw [←inner_conj_sym]; refl, + have : ⟪y, x⟫_ℝ = ⟪x, y⟫_ℝ := by rw [←inner_conj_symm]; refl, simp only [inner_add_add_self, this, add_left_inj], ring, end @@ -591,7 +591,7 @@ by simp only [inner_sub_left, inner_sub_right]; ring /-- Expand `⟪x - y, x - y⟫_ℝ` -/ lemma real_inner_sub_sub_self (x y : F) : ⟪x - y, x - y⟫_ℝ = ⟪x, x⟫_ℝ - 2 * ⟪x, y⟫_ℝ + ⟪y, y⟫_ℝ := begin - have : ⟪y, x⟫_ℝ = ⟪x, y⟫_ℝ := by rw [←inner_conj_sym]; refl, + have : ⟪y, x⟫_ℝ = ⟪x, y⟫_ℝ := by rw [←inner_conj_symm]; refl, simp only [inner_sub_sub_self, this, add_left_inj], ring, end @@ -646,7 +646,7 @@ begin ... = re ⟪x, x⟫ - re (T† * ⟪y, x⟫) - re (T * ⟪x, y⟫) + re (T * T† * ⟪y, y⟫) : by simp only [inner_smul_left, inner_smul_right, mul_assoc] ... = re ⟪x, x⟫ - re (⟪x, y⟫ / ⟪y, y⟫ * ⟪y, x⟫) - : by simp only [map_div₀, h₃, inner_conj_sym, sub_add_cancel] + : by simp only [map_div₀, h₃, inner_conj_symm, sub_add_cancel] with field_simps {discharger := tactic.field_simp.ne_zero} ... = re ⟪x, x⟫ - re (⟪x, y⟫ * ⟪y, x⟫ / ⟪y, y⟫) : by rw ←mul_div_right_comm @@ -666,7 +666,7 @@ end /-- Cauchy–Schwarz inequality for real inner products. -/ lemma real_inner_mul_inner_self_le (x y : F) : ⟪x, y⟫_ℝ * ⟪x, y⟫_ℝ ≤ ⟪x, x⟫_ℝ * ⟪y, y⟫_ℝ := begin - have h₁ : ⟪y, x⟫_ℝ = ⟪x, y⟫_ℝ := by rw [←inner_conj_sym]; refl, + have h₁ : ⟪y, x⟫_ℝ = ⟪x, y⟫_ℝ := by rw [←inner_conj_symm]; refl, have h₂ := @inner_mul_inner_self_le ℝ F _ _ x y, dsimp at h₂, have h₃ := abs_mul_abs_self ⟪x, y⟫_ℝ, @@ -772,7 +772,7 @@ hv.inner_right_sum l (finset.mem_univ _) vectors picks out the coefficient of that vector. -/ lemma orthonormal.inner_left_finsupp {v : ι → E} (hv : orthonormal 𝕜 v) (l : ι →₀ 𝕜) (i : ι) : ⟪finsupp.total ι E 𝕜 v l, v i⟫ = conj (l i) := -by rw [← inner_conj_sym, hv.inner_right_finsupp] +by rw [← inner_conj_symm, hv.inner_right_finsupp] /-- The inner product of a linear combination of a set of orthonormal vectors with one of those vectors picks out the coefficient of that vector. -/ @@ -979,7 +979,7 @@ begin repeat {rw [sq, ←inner_self_eq_norm_mul_norm]}, rw [inner_add_add_self, two_mul], simp only [add_assoc, add_left_inj, add_right_inj, add_monoid_hom.map_add], - rw [←inner_conj_sym, conj_re], + rw [←inner_conj_symm, conj_re], end alias norm_add_sq ← norm_add_pow_two @@ -1007,7 +1007,7 @@ begin re (⟪x, x⟫ - ⟪x, y⟫ - ⟪y, x⟫ + ⟪y, y⟫) = re ⟪x, x⟫ - re ⟪x, y⟫ - re ⟪y, x⟫ + re ⟪y, y⟫ : by simp only [map_add, map_sub] ... = -re ⟪y, x⟫ - re ⟪x, y⟫ + re ⟪x, x⟫ + re ⟪y, y⟫ : by ring - ... = -re (⟪x, y⟫†) - re ⟪x, y⟫ + re ⟪x, x⟫ + re ⟪y, y⟫ : by rw [inner_conj_sym] + ... = -re (⟪x, y⟫†) - re ⟪x, y⟫ + re ⟪x, x⟫ + re ⟪y, y⟫ : by rw [inner_conj_symm] ... = -re ⟪x, y⟫ - re ⟪x, y⟫ + re ⟪x, x⟫ + re ⟪y, y⟫ : by rw [conj_re] ... = re ⟪x, x⟫ - 2*re ⟪x, y⟫ + re ⟪y, y⟫ : by ring end @@ -1035,7 +1035,7 @@ begin have : ‖x‖ * ‖y‖ * (‖x‖ * ‖y‖) = (re ⟪x, x⟫) * (re ⟪y, y⟫), simp only [inner_self_eq_norm_mul_norm], ring, rw this, - conv_lhs { congr, skip, rw [inner_abs_conj_sym] }, + conv_lhs { congr, skip, rw [inner_abs_conj_symm] }, exact inner_mul_inner_self_le _ _ end @@ -1531,7 +1531,7 @@ begin exact h2' }, conv at h2 in ⟪r • x, t⟫ { rw [inner_smul_left, ht0, mul_zero] }, symmetry' at h2, - have h₁ : ⟪t, r • x⟫ = 0 := by { rw [inner_smul_right, ←inner_conj_sym, ht0], simp }, + have h₁ : ⟪t, r • x⟫ = 0 := by { rw [inner_smul_right, ←inner_conj_symm, ht0], simp }, rw [add_zero, h₁, add_left_eq_self, add_zero, inner_self_eq_zero] at h2, rw h2 at ht, exact eq_of_sub_eq_zero ht.symm }, @@ -1801,8 +1801,8 @@ begin simp only [norm_nonneg, pow_nonneg], }, rw [norm_sub_sq, sub_add], simp only [inner_product_space.norm_sq_eq_inner, inner_sum], - simp only [sum_inner, two_mul, inner_smul_right, inner_conj_sym, ←mul_assoc, h₂, ←h₃, - inner_conj_sym, add_monoid_hom.map_sum, finset.mul_sum, ←finset.sum_sub_distrib, inner_smul_left, + simp only [sum_inner, two_mul, inner_smul_right, inner_conj_symm, ←mul_assoc, h₂, ←h₃, + inner_conj_symm, add_monoid_hom.map_sum, finset.mul_sum, ←finset.sum_sub_distrib, inner_smul_left, add_sub_cancel'], end @@ -1835,7 +1835,7 @@ instance is_R_or_C.inner_product_space : inner_product_space 𝕜 𝕜 := inner := λ x y, conj x * y, norm_sq_eq_inner := λ x, by { unfold inner, rw [mul_comm, mul_conj, of_real_re, norm_sq_eq_def'] }, - conj_sym := λ x y, by simp only [mul_comm, map_mul, star_ring_end_self_apply], + conj_symm := λ x y, by simp only [mul_comm, map_mul, star_ring_end_self_apply], add_left := λ x y z, by simp only [add_mul, map_add], smul_left := λ x y z, by simp only [mul_assoc, smul_eq_mul, map_mul] } @@ -1847,7 +1847,7 @@ instance is_R_or_C.inner_product_space : inner_product_space 𝕜 𝕜 := instance submodule.inner_product_space (W : submodule 𝕜 E) : inner_product_space 𝕜 W := { to_normed_add_comm_group := submodule.normed_add_comm_group _, inner := λ x y, ⟪(x:E), (y:E)⟫, - conj_sym := λ _ _, inner_conj_sym _ _ , + conj_symm := λ _ _, inner_conj_symm _ _ , norm_sq_eq_inner := λ _, norm_sq_eq_inner _, add_left := λ _ _ _, inner_add_left _ _ _, smul_left := λ _ _ _, inner_smul_left _ _ _, @@ -2103,7 +2103,7 @@ structure. -/ def inner_product_space.is_R_or_C_to_real : inner_product_space ℝ E := { to_normed_add_comm_group := inner_product_space.to_normed_add_comm_group 𝕜, norm_sq_eq_inner := norm_sq_eq_inner, - conj_sym := λ x y, inner_re_symm _ _, + conj_symm := λ x y, inner_re_symm _ _, add_left := λ x y z, by { change re ⟪x + y, z⟫ = re ⟪x, z⟫ + re ⟪y, z⟫, simp only [inner_add_left, map_add] }, @@ -2221,7 +2221,7 @@ lemma submodule.mem_orthogonal (v : E) : v ∈ Kᗮ ↔ ∀ u ∈ K, ⟪u, v⟫ /-- When a vector is in `Kᗮ`, with the inner product the other way round. -/ lemma submodule.mem_orthogonal' (v : E) : v ∈ Kᗮ ↔ ∀ u ∈ K, ⟪v, u⟫ = 0 := -by simp_rw [submodule.mem_orthogonal, inner_eq_zero_sym] +by simp_rw [submodule.mem_orthogonal, inner_eq_zero_symm] variables {K} @@ -2231,7 +2231,7 @@ lemma submodule.inner_right_of_mem_orthogonal {u v : E} (hu : u ∈ K) (hv : v /-- A vector in `Kᗮ` is orthogonal to one in `K`. -/ lemma submodule.inner_left_of_mem_orthogonal {u v : E} (hu : u ∈ K) (hv : v ∈ Kᗮ) : ⟪v, u⟫ = 0 := -by rw [inner_eq_zero_sym]; exact submodule.inner_right_of_mem_orthogonal hu hv +by rw [inner_eq_zero_symm]; exact submodule.inner_right_of_mem_orthogonal hu hv /-- A vector is in `(𝕜 ∙ u)ᗮ` iff it is orthogonal to `u`. -/ lemma submodule.mem_orthogonal_singleton_iff_inner_right {u v : E} : v ∈ (𝕜 ∙ u)ᗮ ↔ ⟪u, v⟫ = 0 := @@ -2245,7 +2245,7 @@ end /-- A vector in `(𝕜 ∙ u)ᗮ` is orthogonal to `u`. -/ lemma submodule.mem_orthogonal_singleton_iff_inner_left {u v : E} : v ∈ (𝕜 ∙ u)ᗮ ↔ ⟪v, u⟫ = 0 := -by rw [submodule.mem_orthogonal_singleton_iff_inner_right, inner_eq_zero_sym] +by rw [submodule.mem_orthogonal_singleton_iff_inner_right, inner_eq_zero_symm] lemma submodule.sub_mem_orthogonal_of_inner_left {x y : E} (h : ∀ (v : K), ⟪x, v⟫ = ⟪y, v⟫) : x - y ∈ Kᗮ := @@ -2414,11 +2414,11 @@ instance : inner_product_space 𝕜 (completion E) := (continuous_norm.pow 2) (continuous_re.comp (continuous.inner continuous_id' continuous_id'))) (λ a, by simp only [norm_coe, inner_coe, inner_self_eq_norm_sq]), - conj_sym := λ x y, completion.induction_on₂ x y + conj_symm := λ x y, completion.induction_on₂ x y (is_closed_eq (continuous_conj.comp (continuous.inner continuous_snd continuous_fst)) (continuous.inner continuous_fst continuous_snd)) - (λ a b, by simp only [inner_coe, inner_conj_sym]), + (λ a b, by simp only [inner_coe, inner_conj_symm]), add_left := λ x y z, completion.induction_on₃ x y z (is_closed_eq (continuous.inner (continuous_fst.add (continuous_fst.comp continuous_snd)) diff --git a/src/analysis/inner_product_space/dual.lean b/src/analysis/inner_product_space/dual.lean index 588dfa1ec9328..16df24f93439b 100644 --- a/src/analysis/inner_product_space/dual.lean +++ b/src/analysis/inner_product_space/dual.lean @@ -74,8 +74,8 @@ begin refine (function.injective.eq_iff continuous_linear_map.coe_injective).mp (basis.ext b _), intro i, simp only [to_dual_map_apply, continuous_linear_map.coe_coe], - rw [←inner_conj_sym], - nth_rewrite_rhs 0 [←inner_conj_sym], + rw [←inner_conj_symm], + nth_rewrite_rhs 0 [←inner_conj_symm], exact congr_arg conj (h i) end @@ -83,8 +83,8 @@ lemma ext_inner_right_basis {ι : Type*} {x y : E} (b : basis ι 𝕜 E) (h : ∀ i : ι, ⟪x, b i⟫ = ⟪y, b i⟫) : x = y := begin refine ext_inner_left_basis b (λ i, _), - rw [←inner_conj_sym], - nth_rewrite_rhs 0 [←inner_conj_sym], + rw [←inner_conj_symm], + nth_rewrite_rhs 0 [←inner_conj_symm], exact congr_arg conj (h i) end diff --git a/src/analysis/inner_product_space/gram_schmidt_ortho.lean b/src/analysis/inner_product_space/gram_schmidt_ortho.lean index 495227a7c37d0..87cb424373cdb 100644 --- a/src/analysis/inner_product_space/gram_schmidt_ortho.lean +++ b/src/analysis/inner_product_space/gram_schmidt_ortho.lean @@ -83,7 +83,7 @@ begin suffices : ∀ a b : ι, a < b → ⟪gram_schmidt 𝕜 f a, gram_schmidt 𝕜 f b⟫ = 0, { cases h₀.lt_or_lt with ha hb, { exact this _ _ ha, }, - { rw inner_eq_zero_sym, + { rw inner_eq_zero_symm, exact this _ _ hb, }, }, clear h₀ a b, intros a b h₀, @@ -101,7 +101,7 @@ begin simp only [mul_eq_zero, div_eq_zero_iff, inner_self_eq_zero], right, cases hia.lt_or_lt with hia₁ hia₂, - { rw inner_eq_zero_sym, + { rw inner_eq_zero_symm, exact ih a h₀ i hia₁ }, { exact ih i (mem_Iio.1 hi) a hia₂ } end diff --git a/src/analysis/inner_product_space/l2_space.lean b/src/analysis/inner_product_space/l2_space.lean index 8be2b758acb8e..8f92d33f2581a 100644 --- a/src/analysis/inner_product_space/l2_space.lean +++ b/src/analysis/inner_product_space/l2_space.lean @@ -117,10 +117,10 @@ instance : inner_product_space 𝕜 (lp G 2) := { norm_num }, { exact summable_inner f f }, end, - conj_sym := λ f g, begin + conj_symm := λ f g, begin calc conj _ = conj ∑' i, ⟪g i, f i⟫ : by congr ... = ∑' i, conj ⟪g i, f i⟫ : is_R_or_C.conj_cle.map_tsum - ... = ∑' i, ⟪f i, g i⟫ : by simp only [inner_conj_sym] + ... = ∑' i, ⟪f i, g i⟫ : by simp only [inner_conj_symm] ... = _ : by congr, end, add_left := λ f₁ f₂ g, begin @@ -160,7 +160,7 @@ begin end lemma inner_single_right (i : ι) (a : G i) (f : lp G 2) : ⟪f, lp.single 2 i a⟫ = ⟪f i, a⟫ := -by simpa [inner_conj_sym] using congr_arg conj (inner_single_left i a f) +by simpa [inner_conj_symm] using congr_arg conj (inner_single_left i a f) end lp diff --git a/src/analysis/inner_product_space/pi_L2.lean b/src/analysis/inner_product_space/pi_L2.lean index 6fc495ad99628..889680d9b8f36 100644 --- a/src/analysis/inner_product_space/pi_L2.lean +++ b/src/analysis/inner_product_space/pi_L2.lean @@ -74,14 +74,14 @@ instance pi_Lp.inner_product_space {ι : Type*} [fintype ι] (f : ι → Type*) inner := λ x y, ∑ i, inner (x i) (y i), norm_sq_eq_inner := λ x, by simp only [pi_Lp.norm_sq_eq_of_L2, add_monoid_hom.map_sum, ← norm_sq_eq_inner, one_div], - conj_sym := + conj_symm := begin intros x y, unfold inner, rw ring_hom.map_sum, apply finset.sum_congr rfl, rintros z -, - apply inner_conj_sym, + apply inner_conj_symm, end, add_left := λ x y z, show ∑ i, inner (x i + y i) (z i) = ∑ i, inner (x i) (z i) + ∑ i, inner (y i) (z i), diff --git a/src/analysis/inner_product_space/projection.lean b/src/analysis/inner_product_space/projection.lean index 49b4476a27ee1..723cc33dc0589 100644 --- a/src/analysis/inner_product_space/projection.lean +++ b/src/analysis/inner_product_space/projection.lean @@ -456,7 +456,7 @@ orthogonal_projection_fn_inner_eq_zero v v - orthogonal_projection K v ∈ Kᗮ := begin intros w hw, - rw inner_eq_zero_sym, + rw inner_eq_zero_symm, exact orthogonal_projection_inner_eq_zero _ _ hw end @@ -556,7 +556,7 @@ begin obtain ⟨c, rfl⟩ := submodule.mem_span_singleton.mp hx, have hv : ↑‖v‖ ^ 2 = ⟪v, v⟫ := by { norm_cast, simp [norm_sq_eq_inner] }, simp [inner_sub_left, inner_smul_left, inner_smul_right, map_div₀, mul_comm, hv, - inner_product_space.conj_sym, hv] } + inner_product_space.conj_symm, hv] } end /-- Formula for orthogonal projection onto a single vector. -/ @@ -719,11 +719,11 @@ begin { obtain ⟨y, hy, z, hz, rfl⟩ := K.exists_sum_mem_mem_orthogonal v, intros hv, have hz' : z = 0, - { have hyz : ⟪z, y⟫ = 0 := by simp [hz y hy, inner_eq_zero_sym], + { have hyz : ⟪z, y⟫ = 0 := by simp [hz y hy, inner_eq_zero_symm], simpa [inner_add_right, hyz] using hv z hz }, simp [hy, hz'] }, { intros hv w hw, - rw inner_eq_zero_sym, + rw inner_eq_zero_symm, exact hw v hv } end @@ -758,7 +758,7 @@ orthogonal projection. -/ lemma eq_orthogonal_projection_of_mem_orthogonal [complete_space K] {u v : E} (hv : v ∈ K) (hvo : u - v ∈ Kᗮ) : (orthogonal_projection K u : E) = v := -eq_orthogonal_projection_fn_of_mem_of_inner_eq_zero hv (λ w, inner_eq_zero_sym.mp ∘ (hvo w)) +eq_orthogonal_projection_fn_of_mem_of_inner_eq_zero hv (λ w, inner_eq_zero_symm.mp ∘ (hvo w)) /-- A point in `K` with the orthogonality property (here characterized in terms of `Kᗮ`) must be the orthogonal projection. -/ @@ -977,7 +977,7 @@ calc ⟪orthogonal_projection K v, u⟫ @[simp] lemma inner_orthogonal_projection_eq_of_mem_left [complete_space K] (u : K) (v : E) : ⟪u, orthogonal_projection K v⟫ = ⟪(u : E), v⟫ := -by rw [← inner_conj_sym, ← inner_conj_sym (u : E), inner_orthogonal_projection_eq_of_mem_right] +by rw [← inner_conj_symm, ← inner_conj_symm (u : E), inner_orthogonal_projection_eq_of_mem_right] /-- The orthogonal projection is self-adjoint. -/ lemma inner_orthogonal_projection_left_eq_right @@ -1224,7 +1224,7 @@ begin intros hbe', apply hab', simp [ha, hbe'] }, - rw inner_eq_zero_sym, + rw inner_eq_zero_symm, simpa [ha] using h_end b hb }, rintros ⟨b, hb'⟩ hab', cases eq_or_mem_of_mem_insert hb' with hb hb, diff --git a/src/analysis/inner_product_space/symmetric.lean b/src/analysis/inner_product_space/symmetric.lean index aa19ff3829314..c78aff51c2f58 100644 --- a/src/analysis/inner_product_space/symmetric.lean +++ b/src/analysis/inner_product_space/symmetric.lean @@ -62,7 +62,7 @@ end real lemma is_symmetric.conj_inner_sym {T : E →ₗ[𝕜] E} (hT : is_symmetric T) (x y : E) : conj ⟪T x, y⟫ = ⟪T y, x⟫ := -by rw [hT x y, inner_conj_sym] +by rw [hT x y, inner_conj_symm] @[simp] lemma is_symmetric.apply_clm {T : E →L[𝕜] E} (hT : is_symmetric (T : E →ₗ[𝕜] E)) (x y : E) : ⟪T x, y⟫ = ⟪x, T y⟫ := @@ -139,7 +139,7 @@ begin { intros hT v, apply is_symmetric.conj_inner_sym hT }, { intros h x y, - nth_rewrite 1 ← inner_conj_sym, + nth_rewrite 1 ← inner_conj_symm, nth_rewrite 1 inner_map_polarization, simp only [star_ring_end_apply, star_div', star_sub, star_add, star_mul], simp only [← star_ring_end_apply], diff --git a/src/analysis/quaternion.lean b/src/analysis/quaternion.lean index 3b8116325036b..85c7306260f1a 100644 --- a/src/analysis/quaternion.lean +++ b/src/analysis/quaternion.lean @@ -44,7 +44,7 @@ lemma inner_def (a b : ℍ) : ⟪a, b⟫ = (a * b.conj).re := rfl noncomputable instance : inner_product_space ℝ ℍ := inner_product_space.of_core { inner := has_inner.inner, - conj_sym := λ x y, by simp [inner_def, mul_comm], + conj_symm := λ x y, by simp [inner_def, mul_comm], nonneg_re := λ x, norm_sq_nonneg, definite := λ x, norm_sq_eq_zero.1, add_left := λ x y z, by simp only [inner_def, add_mul, add_re], diff --git a/src/linear_algebra/matrix/ldl.lean b/src/linear_algebra/matrix/ldl.lean index 4c5ba1efe5525..1f0f2f720dd3b 100644 --- a/src/linear_algebra/matrix/ldl.lean +++ b/src/linear_algebra/matrix/ldl.lean @@ -93,7 +93,7 @@ begin { simp only [LDL.diag, hij, diagonal_apply_ne, ne.def, not_false_iff, mul_mul_apply], rw [conj_transpose, transpose_map, transpose_transpose, dot_product_mul_vec, (LDL.lower_inv_orthogonal hS (λ h : j = i, hij h.symm)).symm, - ← inner_conj_sym, mul_vec_transpose, euclidean_space.inner_eq_star_dot_product, + ← inner_conj_symm, mul_vec_transpose, euclidean_space.inner_eq_star_dot_product, ← is_R_or_C.star_def, ← star_dot_product_star, dot_product_comm, star_star], refl } end diff --git a/src/linear_algebra/matrix/pos_def.lean b/src/linear_algebra/matrix/pos_def.lean index aa243db83e8a8..cb182b91d8f2b 100644 --- a/src/linear_algebra/matrix/pos_def.lean +++ b/src/linear_algebra/matrix/pos_def.lean @@ -144,7 +144,7 @@ noncomputable def inner_product_space.of_matrix {M : matrix n n 𝕜} (hM : M.pos_def) : inner_product_space 𝕜 (n → 𝕜) := inner_product_space.of_core { inner := λ x y, dot_product (star x) (M.mul_vec y), - conj_sym := λ x y, by + conj_symm := λ x y, by rw [star_dot_product, star_ring_end_apply, star_star, star_mul_vec, dot_product_mul_vec, hM.is_hermitian.eq], nonneg_re := λ x, diff --git a/src/measure_theory/function/l2_space.lean b/src/measure_theory/function/l2_space.lean index 3f9ecafe7dd00..a4158508000d6 100644 --- a/src/measure_theory/function/l2_space.lean +++ b/src/measure_theory/function/l2_space.lean @@ -158,7 +158,7 @@ end instance inner_product_space : inner_product_space 𝕜 (α →₂[μ] E) := { norm_sq_eq_inner := norm_sq_eq_inner', - conj_sym := λ _ _, by simp_rw [inner_def, ← integral_conj, inner_conj_sym], + conj_symm := λ _ _, by simp_rw [inner_def, ← integral_conj, inner_conj_symm], add_left := add_left', smul_left := smul_left', } From 1a313d8bba1bad05faba71a4a4e9742ab5bd9efd Mon Sep 17 00:00:00 2001 From: negiizhao Date: Tue, 14 Mar 2023 18:20:43 +0000 Subject: [PATCH 0626/1291] feat(order/initial_seg): add lemmas about `acc` and `well_founded` (#18527) Co-authored-by: Eric Wieser --- src/order/initial_seg.lean | 31 +++++++++++++++++++++++++++++++ 1 file changed, 31 insertions(+) diff --git a/src/order/initial_seg.lean b/src/order/initial_seg.lean index 4a4a6a3ca57d8..6649efe760bb1 100644 --- a/src/order/initial_seg.lean +++ b/src/order/initial_seg.lean @@ -168,6 +168,13 @@ def le_add (r : α → α → Prop) (s : β → β → Prop) : r ≼i sum.lex r @[simp] theorem le_add_apply (r : α → α → Prop) (s : β → β → Prop) (a) : le_add r s a = sum.inl a := rfl +protected theorem acc (f : r ≼i s) (a : α) : acc r a ↔ acc s (f a) := +⟨begin + refine λ h, acc.rec_on h (λ a _ ha, acc.intro _ (λ b hb, _)), + obtain ⟨a', rfl⟩ := f.init hb, + exact ha _ (f.map_rel_iff.mp hb), + end, f.to_rel_embedding.acc a⟩ + end initial_seg /-! @@ -331,8 +338,32 @@ def of_is_empty (r : α → α → Prop) [is_empty α] {b : β} (H : ∀ b', ¬ @[reducible] def pempty_to_punit : @empty_relation pempty ≺i @empty_relation punit := @of_is_empty _ _ empty_relation _ _ punit.star $ λ x, not_false +protected theorem acc [is_trans β s] (f : r ≺i s) (a : α) : acc r a ↔ acc s (f a) := +(f : r ≼i s).acc a + end principal_seg +/-- +A relation is well-founded iff every principal segment of it is well-founded. + +In this lemma we use `subrel` to indicate its principal segments because it's usually more +convenient to use. +-/ +theorem well_founded_iff_well_founded_subrel {β : Type*} {s : β → β → Prop} [is_trans β s] : + well_founded s ↔ (∀ b, well_founded (subrel s {b' | s b' b})) := +begin + refine ⟨λ wf b, ⟨λ b', ((principal_seg.of_element _ b).acc b').mpr (wf.apply b')⟩, + λ wf, ⟨λ b, acc.intro _ (λ b' hb', _)⟩⟩, + let f := principal_seg.of_element s b, + obtain ⟨b', rfl⟩ := f.down.mp ((principal_seg.of_element_top s b).symm ▸ hb' : s b' f.top), + exact (f.acc b').mp ((wf b).apply b'), +end + +theorem {u} well_founded_iff_principal_seg {β : Type u} {s : β → β → Prop} [is_trans β s] : + well_founded s ↔ (∀ (α : Type u) (r : α → α → Prop) (f : r ≺i s), well_founded r) := +⟨λ wf α r f, rel_hom_class.well_founded f.to_rel_embedding wf, + λ h, well_founded_iff_well_founded_subrel.mpr (λ b, h _ _ (principal_seg.of_element s b))⟩ + /-! ### Properties of initial and principal segments -/ /-- To an initial segment taking values in a well order, one can associate either a principal From 4681620dafca6a7d710f437bd10fb69428ec2209 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 15 Mar 2023 07:28:57 +0000 Subject: [PATCH 0627/1291] fix(analysis/inner_product_space): make type families explicit for `orthogonal_family` and `is_hilbert_sum` (#18584) Pretty much every single use of `orthogonal_family` was unable to infer this argument and so used `@`. `is_hilbert_sum` is changed for consistency. --- src/analysis/inner_product_space/basic.lean | 15 ++++--- .../inner_product_space/l2_space.lean | 40 +++++++++---------- src/analysis/inner_product_space/pi_L2.lean | 16 ++++---- .../inner_product_space/projection.lean | 4 +- .../inner_product_space/spectrum.lean | 5 +-- 5 files changed, 38 insertions(+), 42 deletions(-) diff --git a/src/analysis/inner_product_space/basic.lean b/src/analysis/inner_product_space/basic.lean index 818d75a5cda04..bb2dfcdbe93c5 100644 --- a/src/analysis/inner_product_space/basic.lean +++ b/src/analysis/inner_product_space/basic.lean @@ -1884,15 +1884,15 @@ product space structure on each of the submodules is important -- for example, w their Hilbert sum (`pi_lp V 2`). For example, given an orthonormal set of vectors `v : ι → E`, we have an associated orthogonal family of one-dimensional subspaces of `E`, which it is convenient to be able to discuss using `ι → 𝕜` rather than `Π i : ι, span 𝕜 (v i)`. -/ -def orthogonal_family {G : ι → Type*} [Π i, inner_product_space 𝕜 (G i)] (V : Π i, G i →ₗᵢ[𝕜] E) : +def orthogonal_family (G : ι → Type*) [Π i, inner_product_space 𝕜 (G i)] (V : Π i, G i →ₗᵢ[𝕜] E) : Prop := ∀ ⦃i j⦄, i ≠ j → ∀ v : G i, ∀ w : G j, ⟪V i v, V j w⟫ = 0 variables {𝕜} {G : ι → Type*} [Π i, inner_product_space 𝕜 (G i)] {V : Π i, G i →ₗᵢ[𝕜] E} - (hV : orthogonal_family 𝕜 V) [dec_V : Π i (x : G i), decidable (x ≠ 0)] + (hV : orthogonal_family 𝕜 G V) [dec_V : Π i (x : G i), decidable (x ≠ 0)] lemma orthonormal.orthogonal_family {v : ι → E} (hv : orthonormal 𝕜 v) : - @orthogonal_family 𝕜 _ _ _ _ (λ i : ι, 𝕜) _ + orthogonal_family 𝕜 (λ i : ι, 𝕜) (λ i, linear_isometry.to_span_singleton 𝕜 E (hv.1 i)) := λ i j hij a b, by simp [inner_smul_left, inner_smul_right, hv.2 hij] @@ -1956,7 +1956,7 @@ end /-- The composition of an orthogonal family of subspaces with an injective function is also an orthogonal family. -/ lemma orthogonal_family.comp {γ : Type*} {f : γ → ι} (hf : function.injective f) : - orthogonal_family 𝕜 (λ g : γ, (V (f g) : G (f g) →ₗᵢ[𝕜] E)) := + orthogonal_family 𝕜 (λ g, G (f g)) (λ g, V (f g)) := λ i j hij v w, hV (hf.ne hij) v w lemma orthogonal_family.orthonormal_sigma_orthonormal {α : ι → Type*} {v_family : Π i, (α i) → G i} @@ -2055,7 +2055,7 @@ omit hV elements each from a different subspace in the family is linearly independent. In particular, the pairwise intersections of elements of the family are 0. -/ lemma orthogonal_family.independent {V : ι → submodule 𝕜 E} - (hV : @orthogonal_family 𝕜 _ _ _ _ (λ i, V i) _ (λ i, (V i).subtypeₗᵢ)) : + (hV : orthogonal_family 𝕜 (λ i, V i) (λ i, (V i).subtypeₗᵢ)) : complete_lattice.independent V := begin classical, @@ -2075,7 +2075,7 @@ end include dec_ι lemma direct_sum.is_internal.collected_basis_orthonormal {V : ι → submodule 𝕜 E} - (hV : @orthogonal_family 𝕜 _ _ _ _ (λ i, V i) _ (λ i, (V i).subtypeₗᵢ)) + (hV : orthogonal_family 𝕜 (λ i, V i) (λ i, (V i).subtypeₗᵢ)) (hV_sum : direct_sum.is_internal (λ i, V i)) {α : ι → Type*} {v_family : Π i, basis (α i) 𝕜 (V i)} (hv_family : ∀ i, orthonormal 𝕜 (v_family i)) : @@ -2366,8 +2366,7 @@ begin end lemma submodule.orthogonal_family_self : - @orthogonal_family 𝕜 E _ _ _ (λ b, ((cond b K Kᗮ : submodule 𝕜 E) : Type*)) _ - (λ b, (cond b K Kᗮ).subtypeₗᵢ) + orthogonal_family 𝕜 (λ b, ↥(cond b K Kᗮ)) (λ b, (cond b K Kᗮ).subtypeₗᵢ) | tt tt := absurd rfl | tt ff := λ _ x y, submodule.inner_right_of_mem_orthogonal x.prop y.prop | ff tt := λ _ x y, submodule.inner_left_of_mem_orthogonal y.prop x.prop diff --git a/src/analysis/inner_product_space/l2_space.lean b/src/analysis/inner_product_space/l2_space.lean index 8f92d33f2581a..0b5fbdcc5b84c 100644 --- a/src/analysis/inner_product_space/l2_space.lean +++ b/src/analysis/inner_product_space/l2_space.lean @@ -16,7 +16,7 @@ dependent functions `f : Π i, G i` for which `∑' i, ‖f i‖ ^ 2`, the sum o summable. This construction is sometimes called the *Hilbert sum* of the family `G`. By choosing `G` to be `ι → 𝕜`, the Hilbert space `ℓ²(ι, 𝕜)` may be seen as a special case of this construction. -We also define a *predicate* `is_hilbert_sum 𝕜 E V`, where `V : Π i, G i →ₗᵢ[𝕜] E`, expressing that +We also define a *predicate* `is_hilbert_sum 𝕜 G V`, where `V : Π i, G i →ₗᵢ[𝕜] E`, expressing that `V` is an `orthogonal_family` and that the associated map `lp G 2 →ₗᵢ[𝕜] E` is surjective. ## Main definitions @@ -28,7 +28,7 @@ We also define a *predicate* `is_hilbert_sum 𝕜 E V`, where `V : Π i, G i → * `is_hilbert_sum`: Given a Hilbert space `E`, a family `G` of inner product spaces and a family `V : Π i, G i →ₗᵢ[𝕜] E` of isometric embeddings of the `G i` into `E`, - `is_hilbert_sum 𝕜 E V` means that `V` is an `orthogonal_family` and that the above + `is_hilbert_sum 𝕜 G V` means that `V` is an `orthogonal_family` and that the above linear isometry is surjective. * `is_hilbert_sum.linear_isometry_equiv`: If a Hilbert space `E` is a Hilbert sum of the @@ -167,7 +167,7 @@ end lp /-! ### Identification of a general Hilbert space `E` with a Hilbert sum -/ namespace orthogonal_family -variables {V : Π i, G i →ₗᵢ[𝕜] E} (hV : orthogonal_family 𝕜 V) +variables {V : Π i, G i →ₗᵢ[𝕜] E} (hV : orthogonal_family 𝕜 G V) include cplt hV @@ -253,7 +253,7 @@ end orthogonal_family section is_hilbert_sum -variables (𝕜 E) (V : Π i, G i →ₗᵢ[𝕜] E) (F : ι → submodule 𝕜 E) +variables (𝕜 G) (V : Π i, G i →ₗᵢ[𝕜] E) (F : ι → submodule 𝕜 E) include cplt /-- Given a family of Hilbert spaces `G : ι → Type*`, a Hilbert sum of `G` consists of a Hilbert @@ -263,17 +263,17 @@ space `E` and an orthogonal family `V : Π i, G i →ₗᵢ[𝕜] E` such that t Keeping in mind that `lp G 2` is "the" external Hilbert sum of `G : ι → Type*`, this is analogous to `direct_sum.is_internal`, except that we don't express it in terms of actual submodules. -/ @[protect_proj] structure is_hilbert_sum : Prop := of_surjective :: -(orthogonal_family : orthogonal_family 𝕜 V) +(orthogonal_family : orthogonal_family 𝕜 G V) (surjective_isometry : function.surjective (orthogonal_family.linear_isometry)) -variables {𝕜 E V} +variables {𝕜 G V} /-- If `V : Π i, G i →ₗᵢ[𝕜] E` is an orthogonal family such that the supremum of the ranges of `V i` is dense, then `(E, V)` is a Hilbert sum of `G`. -/ lemma is_hilbert_sum.mk [Π i, complete_space $ G i] - (hVortho : orthogonal_family 𝕜 V) + (hVortho : orthogonal_family 𝕜 G V) (hVtotal : ⊤ ≤ (⨆ i, (V i).to_linear_map.range).topological_closure) : - is_hilbert_sum 𝕜 E V := + is_hilbert_sum 𝕜 G V := { orthogonal_family := hVortho, surjective_isometry := begin @@ -284,16 +284,16 @@ lemma is_hilbert_sum.mk [Π i, complete_space $ G i] /-- This is `orthogonal_family.is_hilbert_sum` in the case of actual inclusions from subspaces. -/ lemma is_hilbert_sum.mk_internal [Π i, complete_space $ F i] - (hFortho : @orthogonal_family 𝕜 E _ _ _ (λ i, F i) _ (λ i, (F i).subtypeₗᵢ)) + (hFortho : orthogonal_family 𝕜 (λ i, F i) (λ i, (F i).subtypeₗᵢ)) (hFtotal : ⊤ ≤ (⨆ i, (F i)).topological_closure) : - @is_hilbert_sum _ 𝕜 _ E _ _ (λ i, F i) _ (λ i, (F i).subtypeₗᵢ) := + is_hilbert_sum 𝕜 (λ i, F i) (λ i, (F i).subtypeₗᵢ) := is_hilbert_sum.mk hFortho (by simpa [subtypeₗᵢ_to_linear_map, range_subtype] using hFtotal) /-- *A* Hilbert sum `(E, V)` of `G` is canonically isomorphic to *the* Hilbert sum of `G`, i.e `lp G 2`. Note that this goes in the opposite direction from `orthogonal_family.linear_isometry`. -/ -noncomputable def is_hilbert_sum.linear_isometry_equiv (hV : is_hilbert_sum 𝕜 E V) : +noncomputable def is_hilbert_sum.linear_isometry_equiv (hV : is_hilbert_sum 𝕜 G V) : E ≃ₗᵢ[𝕜] lp G 2 := linear_isometry_equiv.symm $ linear_isometry_equiv.of_surjective @@ -302,7 +302,7 @@ hV.orthogonal_family.linear_isometry hV.surjective_isometry /-- In the canonical isometric isomorphism between a Hilbert sum `E` of `G` and `lp G 2`, a vector `w : lp G 2` is the image of the infinite sum of the associated elements in `E`. -/ protected lemma is_hilbert_sum.linear_isometry_equiv_symm_apply - (hV : is_hilbert_sum 𝕜 E V) (w : lp G 2) : + (hV : is_hilbert_sum 𝕜 G V) (w : lp G 2) : hV.linear_isometry_equiv.symm w = ∑' i, V i (w i) := by simp [is_hilbert_sum.linear_isometry_equiv, orthogonal_family.linear_isometry_apply] @@ -310,7 +310,7 @@ by simp [is_hilbert_sum.linear_isometry_equiv, orthogonal_family.linear_isometry a vector `w : lp G 2` is the image of the infinite sum of the associated elements in `E`, and this sum indeed converges. -/ protected lemma is_hilbert_sum.has_sum_linear_isometry_equiv_symm - (hV : is_hilbert_sum 𝕜 E V) (w : lp G 2) : + (hV : is_hilbert_sum 𝕜 G V) (w : lp G 2) : has_sum (λ i, V i (w i)) (hV.linear_isometry_equiv.symm w) := by simp [is_hilbert_sum.linear_isometry_equiv, orthogonal_family.has_sum_linear_isometry] @@ -318,7 +318,7 @@ by simp [is_hilbert_sum.linear_isometry_equiv, orthogonal_family.has_sum_linear_ `lp G 2`, an "elementary basis vector" in `lp G 2` supported at `i : ι` is the image of the associated element in `E`. -/ @[simp] protected lemma is_hilbert_sum.linear_isometry_equiv_symm_apply_single - (hV : is_hilbert_sum 𝕜 E V) {i : ι} (x : G i) : + (hV : is_hilbert_sum 𝕜 G V) {i : ι} (x : G i) : hV.linear_isometry_equiv.symm (lp.single 2 i x) = V i x := by simp [is_hilbert_sum.linear_isometry_equiv, orthogonal_family.linear_isometry_apply_single] @@ -326,7 +326,7 @@ by simp [is_hilbert_sum.linear_isometry_equiv, orthogonal_family.linear_isometry `lp G 2`, a finitely-supported vector in `lp G 2` is the image of the associated finite sum of elements of `E`. -/ @[simp] protected lemma is_hilbert_sum.linear_isometry_equiv_symm_apply_dfinsupp_sum_single - (hV : is_hilbert_sum 𝕜 E V) (W₀ : Π₀ (i : ι), G i) : + (hV : is_hilbert_sum 𝕜 G V) (W₀ : Π₀ (i : ι), G i) : hV.linear_isometry_equiv.symm (W₀.sum (lp.single 2)) = (W₀.sum (λ i, V i)) := by simp [is_hilbert_sum.linear_isometry_equiv, orthogonal_family.linear_isometry_apply_dfinsupp_sum_single] @@ -335,7 +335,7 @@ by simp [is_hilbert_sum.linear_isometry_equiv, `lp G 2`, a finitely-supported vector in `lp G 2` is the image of the associated finite sum of elements of `E`. -/ @[simp] protected lemma is_hilbert_sum.linear_isometry_equiv_apply_dfinsupp_sum_single - (hV : is_hilbert_sum 𝕜 E V) (W₀ : Π₀ (i : ι), G i) : + (hV : is_hilbert_sum 𝕜 G V) (W₀ : Π₀ (i : ι), G i) : (hV.linear_isometry_equiv (W₀.sum (λ i, V i)) : Π i, G i) = W₀ := begin rw ← hV.linear_isometry_equiv_symm_apply_dfinsupp_sum_single, @@ -348,8 +348,7 @@ end the family of linear isometries `λ i, λ k, k • v i`. -/ lemma orthonormal.is_hilbert_sum {v : ι → E} (hv : orthonormal 𝕜 v) (hsp : ⊤ ≤ (span 𝕜 (set.range v)).topological_closure) : - @is_hilbert_sum _ 𝕜 _ _ _ _ (λ i : ι, 𝕜) _ - (λ i, linear_isometry.to_span_singleton 𝕜 E (hv.1 i)) := + is_hilbert_sum 𝕜 (λ i : ι, 𝕜) (λ i, linear_isometry.to_span_singleton 𝕜 E (hv.1 i)) := is_hilbert_sum.mk hv.orthogonal_family begin convert hsp, @@ -357,10 +356,9 @@ begin end lemma submodule.is_hilbert_sum_orthogonal (K : submodule 𝕜 E) [hK : complete_space K] : - @is_hilbert_sum _ 𝕜 _ E _ _ (λ b, ((cond b K Kᗮ : submodule 𝕜 E) : Type*)) _ - (λ b, (cond b K Kᗮ).subtypeₗᵢ) := + is_hilbert_sum 𝕜 (λ b, ↥(cond b K Kᗮ)) (λ b, (cond b K Kᗮ).subtypeₗᵢ) := begin - haveI : Π b, complete_space ((cond b K Kᗮ : submodule 𝕜 E) : Type*), + haveI : Π b, complete_space ↥(cond b K Kᗮ), { intro b, cases b; exact orthogonal.complete_space K <|> assumption }, diff --git a/src/analysis/inner_product_space/pi_L2.lean b/src/analysis/inner_product_space/pi_L2.lean index 889680d9b8f36..3b851881f4136 100644 --- a/src/analysis/inner_product_space/pi_L2.lean +++ b/src/analysis/inner_product_space/pi_L2.lean @@ -142,7 +142,7 @@ lemma euclidean_space.inner_eq_star_dot_product (x y : euclidean_space 𝕜 ι) from `E` to `pi_Lp 2` of the subspaces equipped with the `L2` inner product. -/ def direct_sum.is_internal.isometry_L2_of_orthogonal_family [decidable_eq ι] {V : ι → submodule 𝕜 E} (hV : direct_sum.is_internal V) - (hV' : @orthogonal_family 𝕜 _ _ _ _ (λ i, V i) _ (λ i, (V i).subtypeₗᵢ)) : + (hV' : orthogonal_family 𝕜 (λ i, V i) (λ i, (V i).subtypeₗᵢ)) : E ≃ₗᵢ[𝕜] pi_Lp 2 (λ i, V i) := begin let e₁ := direct_sum.linear_equiv_fun_on_fintype 𝕜 ι (λ i, V i), @@ -160,7 +160,7 @@ end @[simp] lemma direct_sum.is_internal.isometry_L2_of_orthogonal_family_symm_apply [decidable_eq ι] {V : ι → submodule 𝕜 E} (hV : direct_sum.is_internal V) - (hV' : @orthogonal_family 𝕜 _ _ _ _ (λ i, V i) _ (λ i, (V i).subtypeₗᵢ)) + (hV' : orthogonal_family 𝕜 (λ i, V i) (λ i, (V i).subtypeₗᵢ)) (w : pi_Lp 2 (λ i, V i)) : (hV.isometry_L2_of_orthogonal_family hV').symm w = ∑ i, (w i : E) := begin @@ -592,7 +592,7 @@ variables {A : ι → submodule 𝕜 E} of the components of the direct sum, the disjoint union of these orthonormal bases is an orthonormal basis for `M`. -/ noncomputable def direct_sum.is_internal.collected_orthonormal_basis - (hV : @orthogonal_family 𝕜 _ _ _ _ (λ i, A i) _ (λ i, (A i).subtypeₗᵢ)) + (hV : orthogonal_family 𝕜 (λ i, A i) (λ i, (A i).subtypeₗᵢ)) [decidable_eq ι] (hV_sum : direct_sum.is_internal (λ i, A i)) {α : ι → Type*} [Π i, fintype (α i)] (v_family : Π i, orthonormal_basis (α i) 𝕜 (A i)) : orthonormal_basis (Σ i, α i) 𝕜 E := @@ -602,7 +602,7 @@ by simpa using hV.orthonormal_sigma_orthonormal lemma direct_sum.is_internal.collected_orthonormal_basis_mem [decidable_eq ι] (h : direct_sum.is_internal A) {α : ι → Type*} - [Π i, fintype (α i)] (hV : @orthogonal_family 𝕜 _ _ _ _ (λ i, A i) _ (λ i, (A i).subtypeₗᵢ)) + [Π i, fintype (α i)] (hV : orthogonal_family 𝕜 (λ i, A i) (λ i, (A i).subtypeₗᵢ)) (v : Π i, orthonormal_basis (α i) 𝕜 (A i)) (a : Σ i, α i) : h.collected_orthonormal_basis hV v a ∈ A a.1 := by simp [direct_sum.is_internal.collected_orthonormal_basis] @@ -687,7 +687,7 @@ variables {n : ℕ} (hn : finrank 𝕜 E = n) [decidable_eq ι] /-- Exhibit a bijection between `fin n` and the index set of a certain basis of an `n`-dimensional inner product space `E`. This should not be accessed directly, but only via the subsequent API. -/ @[irreducible] def direct_sum.is_internal.sigma_orthonormal_basis_index_equiv - (hV' : @orthogonal_family 𝕜 _ _ _ _ (λ i, V i) _ (λ i, (V i).subtypeₗᵢ)) : + (hV' : orthogonal_family 𝕜 (λ i, V i) (λ i, (V i).subtypeₗᵢ)) : (Σ i, fin (finrank 𝕜 (V i))) ≃ fin n := let b := hV.collected_orthonormal_basis hV' (λ i, (std_orthonormal_basis 𝕜 (V i))) in fintype.equiv_fin_of_card_eq $ (finite_dimensional.finrank_eq_card_basis b.to_basis).symm.trans hn @@ -695,7 +695,7 @@ fintype.equiv_fin_of_card_eq $ (finite_dimensional.finrank_eq_card_basis b.to_ba /-- An `n`-dimensional `inner_product_space` equipped with a decomposition as an internal direct sum has an orthonormal basis indexed by `fin n` and subordinate to that direct sum. -/ @[irreducible] def direct_sum.is_internal.subordinate_orthonormal_basis - (hV' : @orthogonal_family 𝕜 _ _ _ _ (λ i, V i) _ (λ i, (V i).subtypeₗᵢ)) : + (hV' : orthogonal_family 𝕜 (λ i, V i) (λ i, (V i).subtypeₗᵢ)) : orthonormal_basis (fin n) 𝕜 E := ((hV.collected_orthonormal_basis hV' (λ i, (std_orthonormal_basis 𝕜 (V i)))).reindex (hV.sigma_orthonormal_basis_index_equiv hn hV')) @@ -704,13 +704,13 @@ sum has an orthonormal basis indexed by `fin n` and subordinate to that direct s sum has an orthonormal basis indexed by `fin n` and subordinate to that direct sum. This function provides the mapping by which it is subordinate. -/ @[irreducible] def direct_sum.is_internal.subordinate_orthonormal_basis_index - (a : fin n) (hV' : @orthogonal_family 𝕜 _ _ _ _ (λ i, V i) _ (λ i, (V i).subtypeₗᵢ)) : ι := + (a : fin n) (hV' : orthogonal_family 𝕜 (λ i, V i) (λ i, (V i).subtypeₗᵢ)) : ι := ((hV.sigma_orthonormal_basis_index_equiv hn hV').symm a).1 /-- The basis constructed in `orthogonal_family.subordinate_orthonormal_basis` is subordinate to the `orthogonal_family` in question. -/ lemma direct_sum.is_internal.subordinate_orthonormal_basis_subordinate - (a : fin n) (hV' : @orthogonal_family 𝕜 _ _ _ _ (λ i, V i) _ (λ i, (V i).subtypeₗᵢ)) : + (a : fin n) (hV' : orthogonal_family 𝕜 (λ i, V i) (λ i, (V i).subtypeₗᵢ)) : (hV.subordinate_orthonormal_basis hn hV' a) ∈ V (hV.subordinate_orthonormal_basis_index hn a hV') := by simpa only [direct_sum.is_internal.subordinate_orthonormal_basis, diff --git a/src/analysis/inner_product_space/projection.lean b/src/analysis/inner_product_space/projection.lean index 723cc33dc0589..ddda82f217ec8 100644 --- a/src/analysis/inner_product_space/projection.lean +++ b/src/analysis/inner_product_space/projection.lean @@ -1155,7 +1155,7 @@ variables {ι : Type*} they provide an internal direct sum decomposition of `E`) if and only if their span has trivial orthogonal complement. -/ lemma orthogonal_family.is_internal_iff_of_is_complete [decidable_eq ι] - {V : ι → submodule 𝕜 E} (hV : @orthogonal_family 𝕜 _ _ _ _ (λ i, V i) _ (λ i, (V i).subtypeₗᵢ)) + {V : ι → submodule 𝕜 E} (hV : orthogonal_family 𝕜 (λ i, V i) (λ i, (V i).subtypeₗᵢ)) (hc : is_complete (↑(supr V) : set E)) : direct_sum.is_internal V ↔ (supr V)ᗮ = ⊥ := begin @@ -1168,7 +1168,7 @@ end they provide an internal direct sum decomposition of `E`) if and only if their span has trivial orthogonal complement. -/ lemma orthogonal_family.is_internal_iff [decidable_eq ι] [finite_dimensional 𝕜 E] - {V : ι → submodule 𝕜 E} (hV : @orthogonal_family 𝕜 _ _ _ _ (λ i, V i) _ (λ i, (V i).subtypeₗᵢ)) : + {V : ι → submodule 𝕜 E} (hV : orthogonal_family 𝕜 (λ i, V i) (λ i, (V i).subtypeₗᵢ)) : direct_sum.is_internal V ↔ (supr V)ᗮ = ⊥ := begin haveI h := finite_dimensional.proper_is_R_or_C 𝕜 ↥(supr V), diff --git a/src/analysis/inner_product_space/spectrum.lean b/src/analysis/inner_product_space/spectrum.lean index 0eb8ca36d17ac..631f947432a04 100644 --- a/src/analysis/inner_product_space/spectrum.lean +++ b/src/analysis/inner_product_space/spectrum.lean @@ -76,7 +76,7 @@ end /-- The eigenspaces of a self-adjoint operator are mutually orthogonal. -/ lemma orthogonal_family_eigenspaces : - @orthogonal_family 𝕜 _ _ _ _ (λ μ, eigenspace T μ) _ (λ μ, (eigenspace T μ).subtypeₗᵢ) := + orthogonal_family 𝕜 (λ μ, eigenspace T μ) (λ μ, (eigenspace T μ).subtypeₗᵢ) := begin rintros μ ν hμν ⟨v, hv⟩ ⟨w, hw⟩, by_cases hv' : v = 0, @@ -88,8 +88,7 @@ begin end lemma orthogonal_family_eigenspaces' : - @orthogonal_family 𝕜 _ _ _ _ (λ μ : eigenvalues T, eigenspace T μ) _ - (λ μ, (eigenspace T μ).subtypeₗᵢ) := + orthogonal_family 𝕜 (λ μ : eigenvalues T, eigenspace T μ) (λ μ, (eigenspace T μ).subtypeₗᵢ) := hT.orthogonal_family_eigenspaces.comp subtype.coe_injective /-- The mutual orthogonal complement of the eigenspaces of a self-adjoint operator on an inner From 10bf4f825ad729c5653adc039dafa3622e7f93c9 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Wed, 15 Mar 2023 08:58:05 +0000 Subject: [PATCH 0628/1291] chore(*): add mathlib4 synchronization comments (#18578) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.char_p.basic` * `algebra.char_p.invertible` * `algebra.char_p.pi` * `algebra.char_p.subring` * `algebra.char_p.two` * `algebra.polynomial.group_ring_action` * `analysis.normed.field.basic` * `analysis.normed.field.infinite_sum` * `analysis.normed.field.unit_ball` * `analysis.normed.order.lattice` * `analysis.normed_space.int` * `category_theory.bicategory.locally_discrete` * `category_theory.category.Cat.limit` * `category_theory.category.Pointed` * `category_theory.category.Quiv` * `category_theory.connected_components` * `category_theory.limits.colimit_limit` * `category_theory.limits.connected` * `category_theory.limits.constructions.filtered` * `category_theory.limits.constructions.over.connected` * `category_theory.limits.constructions.zero_objects` * `category_theory.limits.final` * `category_theory.limits.pi` * `category_theory.limits.preserves.opposites` * `category_theory.limits.preserves.shapes.images` * `category_theory.limits.preserves.shapes.kernels` * `category_theory.limits.shapes.comm_sq` * `category_theory.linear.linear_functor` * `category_theory.preadditive.left_exact` * `combinatorics.catalan` * `data.polynomial.algebra_map` * `data.polynomial.hasse_deriv` * `data.polynomial.integral_normalization` * `data.polynomial.lifts` * `data.polynomial.taylor` * `group_theory.free_abelian_group` * `linear_algebra.smodeq` * `ring_theory.localization.integer` * `ring_theory.polynomial.scale_roots` * `ring_theory.polynomial.tower` * `topology.homotopy.basic` * `topology.instances.ereal` * `topology.metric_space.cau_seq_filter` * `topology.semicontinuous` --- src/algebra/char_p/basic.lean | 3 +++ src/algebra/char_p/invertible.lean | 3 +++ src/algebra/char_p/pi.lean | 3 +++ src/algebra/char_p/subring.lean | 3 +++ src/algebra/char_p/two.lean | 3 +++ src/algebra/polynomial/group_ring_action.lean | 3 +++ src/analysis/normed/field/basic.lean | 3 +++ src/analysis/normed/field/infinite_sum.lean | 3 +++ src/analysis/normed/field/unit_ball.lean | 3 +++ src/analysis/normed/order/lattice.lean | 3 +++ src/analysis/normed_space/int.lean | 3 +++ src/category_theory/bicategory/locally_discrete.lean | 3 +++ src/category_theory/category/Cat/limit.lean | 3 +++ src/category_theory/category/Pointed.lean | 3 +++ src/category_theory/category/Quiv.lean | 3 +++ src/category_theory/connected_components.lean | 3 +++ src/category_theory/limits/colimit_limit.lean | 3 +++ src/category_theory/limits/connected.lean | 3 +++ src/category_theory/limits/constructions/filtered.lean | 3 +++ src/category_theory/limits/constructions/over/connected.lean | 3 +++ src/category_theory/limits/constructions/zero_objects.lean | 3 +++ src/category_theory/limits/final.lean | 3 +++ src/category_theory/limits/pi.lean | 3 +++ src/category_theory/limits/preserves/opposites.lean | 3 +++ src/category_theory/limits/preserves/shapes/images.lean | 3 +++ src/category_theory/limits/preserves/shapes/kernels.lean | 3 +++ src/category_theory/limits/shapes/comm_sq.lean | 3 +++ src/category_theory/linear/linear_functor.lean | 3 +++ src/category_theory/preadditive/left_exact.lean | 3 +++ src/combinatorics/catalan.lean | 3 +++ src/data/polynomial/algebra_map.lean | 3 +++ src/data/polynomial/hasse_deriv.lean | 3 +++ src/data/polynomial/integral_normalization.lean | 3 +++ src/data/polynomial/lifts.lean | 3 +++ src/data/polynomial/taylor.lean | 3 +++ src/group_theory/free_abelian_group.lean | 3 +++ src/linear_algebra/smodeq.lean | 3 +++ src/ring_theory/localization/integer.lean | 3 +++ src/ring_theory/polynomial/scale_roots.lean | 3 +++ src/ring_theory/polynomial/tower.lean | 3 +++ src/topology/homotopy/basic.lean | 3 +++ src/topology/instances/ereal.lean | 3 +++ src/topology/metric_space/cau_seq_filter.lean | 3 +++ src/topology/semicontinuous.lean | 3 +++ 44 files changed, 132 insertions(+) diff --git a/src/algebra/char_p/basic.lean b/src/algebra/char_p/basic.lean index 2346c09251dad..dae05cea243f2 100644 --- a/src/algebra/char_p/basic.lean +++ b/src/algebra/char_p/basic.lean @@ -10,6 +10,9 @@ import ring_theory.nilpotent /-! # Characteristic of semirings + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ universes u v diff --git a/src/algebra/char_p/invertible.lean b/src/algebra/char_p/invertible.lean index 1c0259d9519ae..4b265c9f7f510 100644 --- a/src/algebra/char_p/invertible.lean +++ b/src/algebra/char_p/invertible.lean @@ -9,6 +9,9 @@ import algebra.char_p.basic /-! # Invertibility of elements given a characteristic +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file includes some instances of `invertible` for specific numbers in characteristic zero. Some more cases are given as a `def`, to be included only when needed. To construct instances for concrete numbers, diff --git a/src/algebra/char_p/pi.lean b/src/algebra/char_p/pi.lean index 99db45b67779c..b3591a350372b 100644 --- a/src/algebra/char_p/pi.lean +++ b/src/algebra/char_p/pi.lean @@ -9,6 +9,9 @@ import algebra.ring.pi /-! # Characteristic of semirings of functions + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ universes u v diff --git a/src/algebra/char_p/subring.lean b/src/algebra/char_p/subring.lean index 172058f93ef13..7cfaf15dbf084 100644 --- a/src/algebra/char_p/subring.lean +++ b/src/algebra/char_p/subring.lean @@ -9,6 +9,9 @@ import ring_theory.subring.basic /-! # Characteristic of subrings + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ universes u v diff --git a/src/algebra/char_p/two.lean b/src/algebra/char_p/two.lean index 639cc2e8cbc87..c69cce8d5f6c3 100644 --- a/src/algebra/char_p/two.lean +++ b/src/algebra/char_p/two.lean @@ -8,6 +8,9 @@ import algebra.char_p.basic /-! # Lemmas about rings of characteristic two +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains results about `char_p R 2`, in the `char_two` namespace. The lemmas in this file with a `_sq` suffix are just special cases of the `_pow_char` lemmas diff --git a/src/algebra/polynomial/group_ring_action.lean b/src/algebra/polynomial/group_ring_action.lean index 35c3dcb56b2c8..93b34e2dc6752 100644 --- a/src/algebra/polynomial/group_ring_action.lean +++ b/src/algebra/polynomial/group_ring_action.lean @@ -13,6 +13,9 @@ import group_theory.group_action.quotient /-! # Group action on rings applied to polynomials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains instances and definitions relating `mul_semiring_action` to `polynomial`. -/ diff --git a/src/analysis/normed/field/basic.lean b/src/analysis/normed/field/basic.lean index 20aa3ffdb0612..bd9a12ae695e5 100644 --- a/src/analysis/normed/field/basic.lean +++ b/src/analysis/normed/field/basic.lean @@ -10,6 +10,9 @@ import topology.instances.ennreal /-! # Normed fields +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define (semi)normed rings and fields. We also prove some theorems about these definitions. -/ diff --git a/src/analysis/normed/field/infinite_sum.lean b/src/analysis/normed/field/infinite_sum.lean index eb12b5e6bd2e4..67888a205379c 100644 --- a/src/analysis/normed/field/infinite_sum.lean +++ b/src/analysis/normed/field/infinite_sum.lean @@ -8,6 +8,9 @@ import analysis.normed.group.infinite_sum /-! # Multiplying two infinite sums in a normed ring +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we prove various results about `(∑' x : ι, f x) * (∑' y : ι', g y)` in a normed ring. There are similar results proven in `topology/algebra/infinite_sum` (e.g `tsum_mul_tsum`), but in a normed ring we get summability results which aren't true in general. diff --git a/src/analysis/normed/field/unit_ball.lean b/src/analysis/normed/field/unit_ball.lean index 9de3087c3a72d..397e10d308943 100644 --- a/src/analysis/normed/field/unit_ball.lean +++ b/src/analysis/normed/field/unit_ball.lean @@ -9,6 +9,9 @@ import analysis.normed.group.ball_sphere /-! # Algebraic structures on unit balls and spheres +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define algebraic structures (`semigroup`, `comm_semigroup`, `monoid`, `comm_monoid`, `group`, `comm_group`) on `metric.ball (0 : 𝕜) 1`, `metric.closed_ball (0 : 𝕜) 1`, and `metric.sphere (0 : 𝕜) 1`. In each case we use the weakest possible typeclass assumption on `𝕜`, diff --git a/src/analysis/normed/order/lattice.lean b/src/analysis/normed/order/lattice.lean index e9e8318dbba71..6ff7502e02a51 100644 --- a/src/analysis/normed/order/lattice.lean +++ b/src/analysis/normed/order/lattice.lean @@ -10,6 +10,9 @@ import algebra.order.lattice_group /-! # Normed lattice ordered groups +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Motivated by the theory of Banach Lattices, we then define `normed_lattice_add_comm_group` as a lattice with a covariant normed group addition satisfying the solid axiom. diff --git a/src/analysis/normed_space/int.lean b/src/analysis/normed_space/int.lean index 51d30e87fa72d..51b44ac060cd4 100644 --- a/src/analysis/normed_space/int.lean +++ b/src/analysis/normed_space/int.lean @@ -8,6 +8,9 @@ import analysis.normed.field.basic /-! # The integers as normed ring +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains basic facts about the integers as normed ring. Recall that `‖n‖` denotes the norm of `n` as real number. diff --git a/src/category_theory/bicategory/locally_discrete.lean b/src/category_theory/bicategory/locally_discrete.lean index 650a1097b41d5..e7900aab3159b 100644 --- a/src/category_theory/bicategory/locally_discrete.lean +++ b/src/category_theory/bicategory/locally_discrete.lean @@ -10,6 +10,9 @@ import category_theory.bicategory.strict /-! # Locally discrete bicategories +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A category `C` can be promoted to a strict bicategory `locally_discrete C`. The objects and the 1-morphisms in `locally_discrete C` are the same as the objects and the morphisms, respectively, in `C`, and the 2-morphisms in `locally_discrete C` are the equalities between 1-morphisms. In diff --git a/src/category_theory/category/Cat/limit.lean b/src/category_theory/category/Cat/limit.lean index 778299a632e33..d5ece82aebf5c 100644 --- a/src/category_theory/category/Cat/limit.lean +++ b/src/category_theory/category/Cat/limit.lean @@ -10,6 +10,9 @@ import category_theory.limits.preserves.basic /-! # The category of small categories has all small limits. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + An object in the limit consists of a family of objects, which are carried to one another by the functors in the diagram. A morphism between two such objects is a family of morphisms between the corresponding objects, diff --git a/src/category_theory/category/Pointed.lean b/src/category_theory/category/Pointed.lean index e2966bb49e4f2..6d90709014f79 100644 --- a/src/category_theory/category/Pointed.lean +++ b/src/category_theory/category/Pointed.lean @@ -8,6 +8,9 @@ import category_theory.concrete_category.basic /-! # The category of pointed types +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This defines `Pointed`, the category of pointed types. ## TODO diff --git a/src/category_theory/category/Quiv.lean b/src/category_theory/category/Quiv.lean index 2bcf0b8f18308..b01b516fd8fcc 100644 --- a/src/category_theory/category/Quiv.lean +++ b/src/category_theory/category/Quiv.lean @@ -10,6 +10,9 @@ import category_theory.path_category /-! # The category of quivers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The category of (bundled) quivers, and the free/forgetful adjunction between `Cat` and `Quiv`. -/ diff --git a/src/category_theory/connected_components.lean b/src/category_theory/connected_components.lean index 718dd4d8f9508..a04a891f964cf 100644 --- a/src/category_theory/connected_components.lean +++ b/src/category_theory/connected_components.lean @@ -11,6 +11,9 @@ import category_theory.full_subcategory /-! # Connected components of a category +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Defines a type `connected_components J` indexing the connected components of a category, and the full subcategories giving each connected component: `component j : Type u₁`. We show that each `component j` is in fact connected. diff --git a/src/category_theory/limits/colimit_limit.lean b/src/category_theory/limits/colimit_limit.lean index 50e24cd03ff4f..8a2ab702d589f 100644 --- a/src/category_theory/limits/colimit_limit.lean +++ b/src/category_theory/limits/colimit_limit.lean @@ -10,6 +10,9 @@ import category_theory.limits.functor_category /-! # The morphism comparing a colimit of limits with the corresponding limit of colimits. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + For `F : J × K ⥤ C` there is always a morphism $\colim_k \lim_j F(j,k) → \lim_j \colim_k F(j, k)$. While it is not usually an isomorphism, with additional hypotheses on `J` and `K` it may be, in which case we say that "colimits commute with limits". diff --git a/src/category_theory/limits/connected.lean b/src/category_theory/limits/connected.lean index c0f3710678c42..f066ce70b7849 100644 --- a/src/category_theory/limits/connected.lean +++ b/src/category_theory/limits/connected.lean @@ -12,6 +12,9 @@ import category_theory.limits.preserves.basic /-! # Connected limits +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A connected limit is a limit whose shape is a connected category. We give examples of connected categories, and prove that the functor given diff --git a/src/category_theory/limits/constructions/filtered.lean b/src/category_theory/limits/constructions/filtered.lean index 1836ccd9793c5..42c8ea4d84801 100644 --- a/src/category_theory/limits/constructions/filtered.lean +++ b/src/category_theory/limits/constructions/filtered.lean @@ -9,6 +9,9 @@ import category_theory.limits.opposites /-! # Constructing colimits from finite colimits and filtered colimits +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We construct colimits of size `w` from finite colimits and filtered colimits of size `w`. Since `w`-sized colimits are constructured from coequalizers and `w`-sized coproducts, it suffices to construct `w`-sized coproducts from finite coproducts and `w`-sized filtered colimits. diff --git a/src/category_theory/limits/constructions/over/connected.lean b/src/category_theory/limits/constructions/over/connected.lean index 8852f4b275d68..c09beeb8112d4 100644 --- a/src/category_theory/limits/constructions/over/connected.lean +++ b/src/category_theory/limits/constructions/over/connected.lean @@ -10,6 +10,9 @@ import category_theory.is_connected /-! # Connected limits in the over category +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Shows that the forgetful functor `over B ⥤ C` creates connected limits, in particular `over B` has any connected limit which `C` has. -/ diff --git a/src/category_theory/limits/constructions/zero_objects.lean b/src/category_theory/limits/constructions/zero_objects.lean index c33665cd16695..a4d2f763aba75 100644 --- a/src/category_theory/limits/constructions/zero_objects.lean +++ b/src/category_theory/limits/constructions/zero_objects.lean @@ -10,6 +10,9 @@ import category_theory.limits.constructions.binary_products /-! # Limits involving zero objects +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Binary products and coproducts with a zero object always exist, and pullbacks/pushouts over a zero object are products/coproducts. -/ diff --git a/src/category_theory/limits/final.lean b/src/category_theory/limits/final.lean index 18698aca45cd5..f85798fef3d3f 100644 --- a/src/category_theory/limits/final.lean +++ b/src/category_theory/limits/final.lean @@ -12,6 +12,9 @@ import category_theory.limits.types /-! # Final and initial functors +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A functor `F : C ⥤ D` is final if for every `d : D`, the comma category of morphisms `d ⟶ F.obj c` is connected. diff --git a/src/category_theory/limits/pi.lean b/src/category_theory/limits/pi.lean index 870d341575665..f332f4f13ec6e 100644 --- a/src/category_theory/limits/pi.lean +++ b/src/category_theory/limits/pi.lean @@ -9,6 +9,9 @@ import category_theory.limits.has_limits /-! # Limits in the category of indexed families of objects. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given a functor `F : J ⥤ Π i, C i` into a category of indexed families, 1. we can assemble a collection of cones over `F ⋙ pi.eval C i` into a cone over `F` 2. if all those cones are limit cones, the assembled cone is a limit cone, and diff --git a/src/category_theory/limits/preserves/opposites.lean b/src/category_theory/limits/preserves/opposites.lean index 115f22c342131..915b9945217fc 100644 --- a/src/category_theory/limits/preserves/opposites.lean +++ b/src/category_theory/limits/preserves/opposites.lean @@ -9,6 +9,9 @@ import category_theory.limits.preserves.finite /-! # Limit preservation properties of `functor.op` and related constructions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We formulate conditions about `F` which imply that `F.op`, `F.unop`, `F.left_op` and `F.right_op` preserve certain (co)limits. diff --git a/src/category_theory/limits/preserves/shapes/images.lean b/src/category_theory/limits/preserves/shapes/images.lean index 3c0ac14d8d7c2..d88487ddc47a4 100644 --- a/src/category_theory/limits/preserves/shapes/images.lean +++ b/src/category_theory/limits/preserves/shapes/images.lean @@ -9,6 +9,9 @@ import category_theory.limits.constructions.epi_mono /-! # Preserving images +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we show that if a functor preserves span and cospan, then it preserves images. -/ diff --git a/src/category_theory/limits/preserves/shapes/kernels.lean b/src/category_theory/limits/preserves/shapes/kernels.lean index a174f04aa2e1c..7f916b5db6c08 100644 --- a/src/category_theory/limits/preserves/shapes/kernels.lean +++ b/src/category_theory/limits/preserves/shapes/kernels.lean @@ -9,6 +9,9 @@ import category_theory.limits.preserves.shapes.zero /-! # Preserving (co)kernels +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Constructions to relate the notions of preserving (co)kernels and reflecting (co)kernels to concrete (co)forks. diff --git a/src/category_theory/limits/shapes/comm_sq.lean b/src/category_theory/limits/shapes/comm_sq.lean index 9d517995a6ad7..6ff7ca160f1c9 100644 --- a/src/category_theory/limits/shapes/comm_sq.lean +++ b/src/category_theory/limits/shapes/comm_sq.lean @@ -13,6 +13,9 @@ import category_theory.limits.constructions.zero_objects /-! # Pullback and pushout squares, and bicartesian squares +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We provide another API for pullbacks and pushouts. `is_pullback fst snd f g` is the proposition that diff --git a/src/category_theory/linear/linear_functor.lean b/src/category_theory/linear/linear_functor.lean index 6804f4c00d6b7..441907c22cb7e 100644 --- a/src/category_theory/linear/linear_functor.lean +++ b/src/category_theory/linear/linear_functor.lean @@ -9,6 +9,9 @@ import category_theory.linear.basic /-! # Linear Functors +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + An additive functor between two `R`-linear categories is called *linear* if the induced map on hom types is a morphism of `R`-modules. diff --git a/src/category_theory/preadditive/left_exact.lean b/src/category_theory/preadditive/left_exact.lean index d98a6fa81a43f..a2c9a6111c2ca 100644 --- a/src/category_theory/preadditive/left_exact.lean +++ b/src/category_theory/preadditive/left_exact.lean @@ -11,6 +11,9 @@ import category_theory.preadditive.additive_functor /-! # Left exactness of functors between preadditive categories +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We show that a functor is left exact in the sense that it preserves finite limits, if it preserves kernels. The dual result holds for right exact functors and cokernels. ## Main results diff --git a/src/combinatorics/catalan.lean b/src/combinatorics/catalan.lean index b64c70fa5fd44..484d6028787cf 100644 --- a/src/combinatorics/catalan.lean +++ b/src/combinatorics/catalan.lean @@ -15,6 +15,9 @@ import tactic.linear_combination /-! # Catalan numbers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The Catalan numbers (http://oeis.org/A000108) are probably the most ubiquitous sequence of integers in mathematics. They enumerate several important objects like binary trees, Dyck paths, and triangulations of convex polygons. diff --git a/src/data/polynomial/algebra_map.lean b/src/data/polynomial/algebra_map.lean index db400c57c4623..d746a42084401 100644 --- a/src/data/polynomial/algebra_map.lean +++ b/src/data/polynomial/algebra_map.lean @@ -10,6 +10,9 @@ import data.polynomial.eval /-! # Theory of univariate polynomials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We show that `A[X]` is an R-algebra when `A` is an R-algebra. We promote `eval₂` to an algebra hom in `aeval`. -/ diff --git a/src/data/polynomial/hasse_deriv.lean b/src/data/polynomial/hasse_deriv.lean index 33fea89cf0897..6c237af6ae053 100644 --- a/src/data/polynomial/hasse_deriv.lean +++ b/src/data/polynomial/hasse_deriv.lean @@ -12,6 +12,9 @@ import data.polynomial.derivative /-! # Hasse derivative of polynomials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The `k`th Hasse derivative of a polynomial `∑ a_i X^i` is `∑ (i.choose k) a_i X^(i-k)`. It is a variant of the usual derivative, and satisfies `k! * (hasse_deriv k f) = derivative^[k] f`. The main benefit is that is gives an atomic way of talking about expressions such as diff --git a/src/data/polynomial/integral_normalization.lean b/src/data/polynomial/integral_normalization.lean index 21b207a18e9b4..d0cbeceb1a948 100644 --- a/src/data/polynomial/integral_normalization.lean +++ b/src/data/polynomial/integral_normalization.lean @@ -10,6 +10,9 @@ import data.polynomial.monic /-! # Theory of monic polynomials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define `integral_normalization`, which relate arbitrary polynomials to monic ones. -/ diff --git a/src/data/polynomial/lifts.lean b/src/data/polynomial/lifts.lean index deb68c3e51d6c..1b6ea296def3c 100644 --- a/src/data/polynomial/lifts.lean +++ b/src/data/polynomial/lifts.lean @@ -9,6 +9,9 @@ import data.polynomial.monic /-! # Polynomials that lift +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given semirings `R` and `S` with a morphism `f : R →+* S`, we define a subsemiring `lifts` of `S[X]` by the image of `ring_hom.of (map f)`. Then, we prove that a polynomial that lifts can always be lifted to a polynomial of the same degree diff --git a/src/data/polynomial/taylor.lean b/src/data/polynomial/taylor.lean index 5ea229788f0b4..8cea6165daa56 100644 --- a/src/data/polynomial/taylor.lean +++ b/src/data/polynomial/taylor.lean @@ -11,6 +11,9 @@ import data.polynomial.degree.lemmas /-! # Taylor expansions of polynomials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main declarations * `polynomial.taylor`: the Taylor expansion of the polynomial `f` at `r` diff --git a/src/group_theory/free_abelian_group.lean b/src/group_theory/free_abelian_group.lean index cdce658dbf398..d9e4dbc8d6405 100644 --- a/src/group_theory/free_abelian_group.lean +++ b/src/group_theory/free_abelian_group.lean @@ -12,6 +12,9 @@ import algebra.module.basic /-! # Free abelian groups +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The free abelian group on a type `α`, defined as the abelianisation of the free group on `α`. diff --git a/src/linear_algebra/smodeq.lean b/src/linear_algebra/smodeq.lean index edba4d40fa9ee..36269e80e2407 100644 --- a/src/linear_algebra/smodeq.lean +++ b/src/linear_algebra/smodeq.lean @@ -9,6 +9,9 @@ import ring_theory.ideal.quotient /-! # modular equivalence for submodule + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ open submodule diff --git a/src/ring_theory/localization/integer.lean b/src/ring_theory/localization/integer.lean index 512ea6a21df5e..65928aeded7b1 100644 --- a/src/ring_theory/localization/integer.lean +++ b/src/ring_theory/localization/integer.lean @@ -8,6 +8,9 @@ import ring_theory.localization.basic /-! # Integer elements of a localization +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions * `is_localization.is_integer` is a predicate stating that `x : S` is in the image of `R` diff --git a/src/ring_theory/polynomial/scale_roots.lean b/src/ring_theory/polynomial/scale_roots.lean index 393103daa7ccc..81f1f9a01c644 100644 --- a/src/ring_theory/polynomial/scale_roots.lean +++ b/src/ring_theory/polynomial/scale_roots.lean @@ -10,6 +10,9 @@ import data.polynomial.algebra_map /-! # Scaling the roots of a polynomial +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines `scale_roots p s` for a polynomial `p` in one variable and a ring element `s` to be the polynomial with root `r * s` for each root `r` of `p` and proves some basic results about it. -/ diff --git a/src/ring_theory/polynomial/tower.lean b/src/ring_theory/polynomial/tower.lean index 7342032c8b52b..8dc8313ae783e 100644 --- a/src/ring_theory/polynomial/tower.lean +++ b/src/ring_theory/polynomial/tower.lean @@ -10,6 +10,9 @@ import data.polynomial.algebra_map /-! # Algebra towers for polynomial +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves some basic results about the algebra tower structure for the type `R[X]`. This structure itself is provided elsewhere as `polynomial.is_scalar_tower` diff --git a/src/topology/homotopy/basic.lean b/src/topology/homotopy/basic.lean index 7deeee09350bc..cf1216c36bf3a 100644 --- a/src/topology/homotopy/basic.lean +++ b/src/topology/homotopy/basic.lean @@ -12,6 +12,9 @@ import topology.unit_interval /-! # Homotopy between functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we define a homotopy between two functions `f₀` and `f₁`. First we define `continuous_map.homotopy` between the two functions, with no restrictions on the intermediate maps. Then, as in the formalisation in HOL-Analysis, we define diff --git a/src/topology/instances/ereal.lean b/src/topology/instances/ereal.lean index 39238619680bc..6daf580037099 100644 --- a/src/topology/instances/ereal.lean +++ b/src/topology/instances/ereal.lean @@ -11,6 +11,9 @@ import topology.instances.ennreal /-! # Topological structure on `ereal` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We endow `ereal` with the order topology, and prove basic properties of this topology. ## Main results diff --git a/src/topology/metric_space/cau_seq_filter.lean b/src/topology/metric_space/cau_seq_filter.lean index 11f3c9b94fe6d..4acfa430161bc 100644 --- a/src/topology/metric_space/cau_seq_filter.lean +++ b/src/topology/metric_space/cau_seq_filter.lean @@ -8,6 +8,9 @@ import analysis.normed.field.basic /-! # Completeness in terms of `cauchy` filters vs `is_cau_seq` sequences +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we apply `metric.complete_of_cauchy_seq_tendsto` to prove that a `normed_ring` is complete in terms of `cauchy` filter if and only if it is complete in terms of `cau_seq` Cauchy sequences. diff --git a/src/topology/semicontinuous.lean b/src/topology/semicontinuous.lean index 7d8121174b2e1..a00f7b288c64e 100644 --- a/src/topology/semicontinuous.lean +++ b/src/topology/semicontinuous.lean @@ -10,6 +10,9 @@ import topology.instances.ennreal /-! # Semicontinuous maps +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A function `f` from a topological space `α` to an ordered space `β` is lower semicontinuous at a point `x` if, for any `y < f x`, for any `x'` close enough to `x`, one has `f x' > y`. In other words, `f` can jump up, but it can not jump down. From 0111834459f5d7400215223ea95ae38a1265a907 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 15 Mar 2023 11:17:35 +0000 Subject: [PATCH 0629/1291] chore(order/with_bot): lemmas about unbot and untop (#18582) --- src/algebra/order/monoid/with_top.lean | 9 ++++++++- src/algebra/order/ring/with_top.lean | 4 ---- src/order/succ_pred/basic.lean | 10 ++++++++++ src/order/with_bot.lean | 4 ++++ 4 files changed, 22 insertions(+), 5 deletions(-) diff --git a/src/algebra/order/monoid/with_top.lean b/src/algebra/order/monoid/with_top.lean index 3648e90254f39..616d4148549cb 100644 --- a/src/algebra/order/monoid/with_top.lean +++ b/src/algebra/order/monoid/with_top.lean @@ -11,7 +11,8 @@ import data.nat.cast.defs /-! # Adjoining top/bottom elements to ordered monoids. > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> Any changes to this file require a corresponding PR to mathlib4.-/ +> Any changes to this file require a corresponding PR to mathlib4. +-/ universes u v variables {α : Type u} {β : Type v} @@ -31,6 +32,9 @@ variables [has_one α] @[simp, norm_cast, to_additive] lemma coe_eq_one {a : α} : (a : with_top α) = 1 ↔ a = 1 := coe_eq_coe +@[simp, to_additive] lemma untop_one : (1 : with_top α).untop coe_ne_top = 1 := rfl +@[simp, to_additive] lemma untop_one' (d : α) : (1 : with_top α).untop' d = 1 := rfl + @[simp, norm_cast, to_additive coe_nonneg] lemma one_le_coe [has_le α] {a : α} : 1 ≤ (a : with_top α) ↔ 1 ≤ a := coe_le_coe @@ -347,6 +351,9 @@ lemma coe_one [has_one α] : ((1 : α) : with_bot α) = 1 := rfl lemma coe_eq_one [has_one α] {a : α} : (a : with_bot α) = 1 ↔ a = 1 := with_top.coe_eq_one +@[simp, to_additive] lemma unbot_one [has_one α] : (1 : with_bot α).unbot coe_ne_bot = 1 := rfl +@[simp, to_additive] lemma unbot_one' [has_one α] (d : α) : (1 : with_bot α).unbot' d = 1 := rfl + @[simp, norm_cast, to_additive coe_nonneg] lemma one_le_coe [has_one α] [has_le α] {a : α} : 1 ≤ (a : with_bot α) ↔ 1 ≤ a := coe_le_coe diff --git a/src/algebra/order/ring/with_top.lean b/src/algebra/order/ring/with_top.lean index d39db797bd686..aa0716b3ef5dd 100644 --- a/src/algebra/order/ring/with_top.lean +++ b/src/algebra/order/ring/with_top.lean @@ -20,8 +20,6 @@ variables {α : Type*} namespace with_top -instance [nonempty α] : nontrivial (with_top α) := option.nontrivial - variable [decidable_eq α] section has_mul @@ -198,8 +196,6 @@ end with_top namespace with_bot -instance [nonempty α] : nontrivial (with_bot α) := option.nontrivial - variable [decidable_eq α] section has_mul diff --git a/src/order/succ_pred/basic.lean b/src/order/succ_pred/basic.lean index 4b540794fae56..6f0151b672335 100644 --- a/src/order/succ_pred/basic.lean +++ b/src/order/succ_pred/basic.lean @@ -825,6 +825,11 @@ instance : pred_order (with_top α) := @[simp] lemma pred_top : pred (⊤ : with_top α) = ↑(⊤ : α) := rfl @[simp] lemma pred_coe (a : α) : pred (↑a : with_top α) = ↑(pred a) := rfl +@[simp] lemma pred_untop : ∀ (a : with_top α) (ha : a ≠ ⊤), + pred (a.untop ha) = (pred a).untop (by induction a using with_top.rec_top_coe; simp) +| ⊤ ha := (ha rfl).elim +| (a : α) ha := rfl + end pred /-! #### Adding a `⊤` to a `no_max_order` -/ @@ -922,6 +927,11 @@ instance : succ_order (with_bot α) := @[simp] lemma succ_bot : succ (⊥ : with_bot α) = ↑(⊥ : α) := rfl @[simp] lemma succ_coe (a : α) : succ (↑a : with_bot α) = ↑(succ a) := rfl +@[simp] lemma succ_unbot : ∀ (a : with_bot α) (ha : a ≠ ⊥), + succ (a.unbot ha) = (succ a).unbot (by induction a using with_bot.rec_bot_coe; simp) +| ⊥ ha := (ha rfl).elim +| (a : α) ha := rfl + end succ section pred diff --git a/src/order/with_bot.lean b/src/order/with_bot.lean index 2ad1002fe61ca..4ad53041d78d9 100644 --- a/src/order/with_bot.lean +++ b/src/order/with_bot.lean @@ -48,6 +48,8 @@ meta instance {α : Type} [reflected _ α] [has_reflect α] : has_reflect (with_ instance : inhabited (with_bot α) := ⟨⊥⟩ +instance [nonempty α] : nontrivial (with_bot α) := option.nontrivial + open function lemma coe_injective : injective (coe : α → with_bot α) := option.some_injective _ @@ -420,6 +422,8 @@ meta instance {α : Type} [reflected _ α] [has_reflect α] : has_reflect (with_ instance : inhabited (with_top α) := ⟨⊤⟩ +instance [nonempty α] : nontrivial (with_top α) := option.nontrivial + protected lemma «forall» {p : with_top α → Prop} : (∀ x, p x) ↔ p ⊤ ∧ ∀ x : α, p x := option.forall protected lemma «exists» {p : with_top α → Prop} : (∃ x, p x) ↔ p ⊤ ∨ ∃ x : α, p x := option.exists From 29d5700b0872ffc82431073f46c5b1092f74151c Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 15 Mar 2023 11:17:36 +0000 Subject: [PATCH 0630/1291] chore(geometry/manifold/instances/sphere): speedup a proof (#18586) This goes from 15s to 1.75 seconds. I've squeezed this in multiple steps to make the proof easier to follow. --- src/geometry/manifold/instances/sphere.lean | 11 +++++++++-- 1 file changed, 9 insertions(+), 2 deletions(-) diff --git a/src/geometry/manifold/instances/sphere.lean b/src/geometry/manifold/instances/sphere.lean index 1ad1ee1abe08f..45b43f2d94681 100644 --- a/src/geometry/manifold/instances/sphere.lean +++ b/src/geometry/manifold/instances/sphere.lean @@ -398,8 +398,15 @@ begin have H₂ := (cont_diff_stereo_inv_fun_aux.comp (ℝ ∙ (v:E))ᗮ.subtypeL.cont_diff).comp U.symm.cont_diff, convert H₁.comp' (H₂.cont_diff_on : cont_diff_on ℝ ⊤ _ set.univ) using 1, - ext, - simp [sphere_ext_iff, stereographic'_symm_apply, real_inner_comm] + -- squeezed from `ext, simp [sphere_ext_iff, stereographic'_symm_apply, real_inner_comm]` + simp only [local_homeomorph.trans_to_local_equiv, local_homeomorph.symm_to_local_equiv, + local_equiv.trans_source, local_equiv.symm_source, stereographic'_target, + stereographic'_source], + simp only [model_with_corners_self_coe, model_with_corners_self_coe_symm, set.preimage_id, + set.range_id, set.inter_univ, set.univ_inter, set.compl_singleton_eq, set.preimage_set_of_eq], + simp only [id.def, comp_apply, submodule.subtypeL_apply, local_homeomorph.coe_coe_symm, + innerSL_apply, ne.def, sphere_ext_iff, real_inner_comm (v' : E)], + refl, end /-- The inclusion map (i.e., `coe`) from the sphere in `E` to `E` is smooth. -/ From 992efbda6f85a5c9074375d3c7cb9764c64d8f72 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 15 Mar 2023 14:32:20 +0000 Subject: [PATCH 0631/1291] feat(analysis/normed/order/upper_lower): Thickening an upper set (#17257) Results about upper sets in normed ordered groups. --- src/analysis/normed/order/upper_lower.lean | 139 ++++++++++++++++++++ src/topology/algebra/order/upper_lower.lean | 7 +- 2 files changed, 142 insertions(+), 4 deletions(-) create mode 100644 src/analysis/normed/order/upper_lower.lean diff --git a/src/analysis/normed/order/upper_lower.lean b/src/analysis/normed/order/upper_lower.lean new file mode 100644 index 0000000000000..fa1fe8f25bf1a --- /dev/null +++ b/src/analysis/normed/order/upper_lower.lean @@ -0,0 +1,139 @@ +/- +Copyright (c) 2022 Yaël Dillies. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies +-/ +import algebra.order.field.pi +import analysis.normed.group.pointwise +import analysis.normed.order.basic +import topology.algebra.order.upper_lower + +/-! +# Upper/lower/order-connected sets in normed groups + +The topological closure and interior of an upper/lower/order-connected set is an +upper/lower/order-connected set (with the notable exception of the closure of an order-connected +set). + +We also prove lemmas specific to `ℝⁿ`. Those are helpful to prove that order-connected sets in `ℝⁿ` +are measurable. +-/ + +open function metric set + +variables {α ι : Type*} + +section metric_space +variables [normed_ordered_group α] {s : set α} + +@[to_additive is_upper_set.thickening] +protected lemma is_upper_set.thickening' (hs : is_upper_set s) (ε : ℝ) : + is_upper_set (thickening ε s) := +by { rw ←ball_mul_one, exact hs.mul_left } + +@[to_additive is_lower_set.thickening] +protected lemma is_lower_set.thickening' (hs : is_lower_set s) (ε : ℝ) : + is_lower_set (thickening ε s) := +by { rw ←ball_mul_one, exact hs.mul_left } + +@[to_additive is_upper_set.cthickening] +protected lemma is_upper_set.cthickening' (hs : is_upper_set s) (ε : ℝ) : + is_upper_set (cthickening ε s) := +by { rw cthickening_eq_Inter_thickening'', exact is_upper_set_Inter₂ (λ δ hδ, hs.thickening' _) } + +@[to_additive is_lower_set.cthickening] +protected lemma is_lower_set.cthickening' (hs : is_lower_set s) (ε : ℝ) : + is_lower_set (cthickening ε s) := +by { rw cthickening_eq_Inter_thickening'', exact is_lower_set_Inter₂ (λ δ hδ, hs.thickening' _) } + +end metric_space + +/-! ### `ℝⁿ` -/ + +section finite +variables [finite ι] {s : set (ι → ℝ)} {x y : ι → ℝ} {δ : ℝ} + +lemma is_upper_set.mem_interior_of_forall_lt (hs : is_upper_set s) (hx : x ∈ closure s) + (h : ∀ i, x i < y i) : + y ∈ interior s := +begin + casesI nonempty_fintype ι, + obtain ⟨ε, hε, hxy⟩ := pi.exists_forall_pos_add_lt h, + obtain ⟨z, hz, hxz⟩ := metric.mem_closure_iff.1 hx _ hε, + rw dist_pi_lt_iff hε at hxz, + have hyz : ∀ i, z i < y i, + { refine λ i, (hxy _).trans_le' (sub_le_iff_le_add'.1 $ (le_abs_self _).trans _), + rw [←real.norm_eq_abs, ←dist_eq_norm'], + exact (hxz _).le }, + obtain ⟨δ, hδ, hyz⟩ := pi.exists_forall_pos_add_lt hyz, + refine mem_interior.2 ⟨ball y δ, _, is_open_ball, mem_ball_self hδ⟩, + rintro w hw, + refine hs (λ i, _) hz, + simp_rw [ball_pi _ hδ, real.ball_eq_Ioo] at hw, + exact ((lt_sub_iff_add_lt.2 $ hyz _).trans (hw _ $ mem_univ _).1).le, +end + +lemma is_lower_set.mem_interior_of_forall_lt (hs : is_lower_set s) (hx : x ∈ closure s) + (h : ∀ i, y i < x i) : + y ∈ interior s := +begin + casesI nonempty_fintype ι, + obtain ⟨ε, hε, hxy⟩ := pi.exists_forall_pos_add_lt h, + obtain ⟨z, hz, hxz⟩ := metric.mem_closure_iff.1 hx _ hε, + rw dist_pi_lt_iff hε at hxz, + have hyz : ∀ i, y i < z i, + { refine λ i, (lt_sub_iff_add_lt.2 $ hxy _).trans_le (sub_le_comm.1 $ (le_abs_self _).trans _), + rw [←real.norm_eq_abs, ←dist_eq_norm], + exact (hxz _).le }, + obtain ⟨δ, hδ, hyz⟩ := pi.exists_forall_pos_add_lt hyz, + refine mem_interior.2 ⟨ball y δ, _, is_open_ball, mem_ball_self hδ⟩, + rintro w hw, + refine hs (λ i, _) hz, + simp_rw [ball_pi _ hδ, real.ball_eq_Ioo] at hw, + exact ((hw _ $ mem_univ _).2.trans $ hyz _).le, +end + +end finite + +section fintype +variables [fintype ι] {s : set (ι → ℝ)} {x y : ι → ℝ} {δ : ℝ} + +lemma is_upper_set.exists_subset_ball (hs : is_upper_set s) (hx : x ∈ closure s) (hδ : 0 < δ) : + ∃ y, closed_ball y (δ/4) ⊆ closed_ball x δ ∧ closed_ball y (δ/4) ⊆ interior s := +begin + refine ⟨x + const _ (3/4*δ), closed_ball_subset_closed_ball' _, _⟩, + { rw dist_self_add_left, + refine (add_le_add_left (pi_norm_const_le $ 3 / 4 * δ) _).trans_eq _, + simp [real.norm_of_nonneg, hδ.le, zero_le_three], + ring_nf }, + obtain ⟨y, hy, hxy⟩ := metric.mem_closure_iff.1 hx _ (div_pos hδ zero_lt_four), + refine λ z hz, hs.mem_interior_of_forall_lt (subset_closure hy) (λ i, _), + rw [mem_closed_ball, dist_eq_norm'] at hz, + rw dist_eq_norm at hxy, + replace hxy := (norm_le_pi_norm _ i).trans hxy.le, + replace hz := (norm_le_pi_norm _ i).trans hz, + dsimp at hxy hz, + rw abs_sub_le_iff at hxy hz, + linarith, +end + +lemma is_lower_set.exists_subset_ball (hs : is_lower_set s) (hx : x ∈ closure s) (hδ : 0 < δ) : + ∃ y, closed_ball y (δ/4) ⊆ closed_ball x δ ∧ closed_ball y (δ/4) ⊆ interior s := +begin + refine ⟨x - const _ (3/4*δ), closed_ball_subset_closed_ball' _, _⟩, + { rw dist_self_sub_left, + refine (add_le_add_left (pi_norm_const_le $ 3 / 4 * δ) _).trans_eq _, + simp [real.norm_of_nonneg, hδ.le, zero_le_three], + ring_nf }, + obtain ⟨y, hy, hxy⟩ := metric.mem_closure_iff.1 hx _ (div_pos hδ zero_lt_four), + refine λ z hz, hs.mem_interior_of_forall_lt (subset_closure hy) (λ i, _), + rw [mem_closed_ball, dist_eq_norm'] at hz, + rw dist_eq_norm at hxy, + replace hxy := (norm_le_pi_norm _ i).trans hxy.le, + replace hz := (norm_le_pi_norm _ i).trans hz, + dsimp at hxy hz, + rw abs_sub_le_iff at hxy hz, + linarith, +end + +end fintype diff --git a/src/topology/algebra/order/upper_lower.lean b/src/topology/algebra/order/upper_lower.lean index 59c9368d26e5a..5bc75c68c9be3 100644 --- a/src/topology/algebra/order/upper_lower.lean +++ b/src/topology/algebra/order/upper_lower.lean @@ -16,11 +16,10 @@ The topological closure and interior of an upper/lower/order-connected set is an upper/lower/order-connected set (with the notable exception of the closure of an order-connected set). -## Notes +## Implementation notes -The lemmas don't mention additive/multiplicative operations. As a result, we decide to prime the -multiplicative lemma names to indicate that there is probably a common generalisation to each pair -of additive/multiplicative lemma. +The same lemmas are true in the additive/multiplicative worlds. To avoid code duplication, we +provide `has_upper_lower_closure`, an ad hoc axiomatisation of the properties we need. -/ open function set From 83f81aea33931a1edb94ce0f32b9a5d484de6978 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Wed, 15 Mar 2023 14:32:21 +0000 Subject: [PATCH 0632/1291] refactor(topology/algebra/open_subgroup): use `set_like` (#18585) --- src/field_theory/krull_topology.lean | 4 +- .../algebra/nonarchimedean/basic.lean | 5 +- src/topology/algebra/open_subgroup.lean | 163 +++++++++--------- 3 files changed, 83 insertions(+), 89 deletions(-) diff --git a/src/field_theory/krull_topology.lean b/src/field_theory/krull_topology.lean index 03c6438783c48..b24e897daa3cb 100644 --- a/src/field_theory/krull_topology.lean +++ b/src/field_theory/krull_topology.lean @@ -210,9 +210,7 @@ begin have h_basis : E.fixing_subgroup.carrier ∈ (gal_group_basis K L) := ⟨E.fixing_subgroup, ⟨E, ‹_›, rfl⟩, rfl⟩, have h_nhd := group_filter_basis.mem_nhds_one (gal_group_basis K L) h_basis, - rw mem_nhds_iff at h_nhd, - rcases h_nhd with ⟨U, hU_le, hU_open, h1U⟩, - exact subgroup.is_open_of_one_mem_interior ⟨U, ⟨hU_open, hU_le⟩, h1U⟩, + exact subgroup.is_open_of_mem_nhds _ h_nhd end /-- Given a tower of fields `L/E/K`, with `E/K` finite, the subgroup `Gal(L/E) ≤ L ≃ₐ[K] L` is diff --git a/src/topology/algebra/nonarchimedean/basic.lean b/src/topology/algebra/nonarchimedean/basic.lean index c724c8c2e2068..a19b98853f8de 100644 --- a/src/topology/algebra/nonarchimedean/basic.lean +++ b/src/topology/algebra/nonarchimedean/basic.lean @@ -127,14 +127,13 @@ lemma mul_subset (U : open_add_subgroup R) : ∃ V : open_add_subgroup R, (V : set R) * V ⊆ U := let ⟨V, H⟩ := prod_self_subset (is_open.mem_nhds (is_open.preimage continuous_mul U.is_open) begin - simpa only [set.mem_preimage, open_add_subgroup.mem_coe, prod.snd_zero, mul_zero] - using U.zero_mem, + simpa only [set.mem_preimage, set_like.mem_coe, prod.snd_zero, mul_zero] using U.zero_mem, end) in begin use V, rintros v ⟨a, b, ha, hb, hv⟩, have hy := H (set.mk_mem_prod ha hb), - simp only [set.mem_preimage, open_add_subgroup.mem_coe] at hy, + simp only [set.mem_preimage, set_like.mem_coe] at hy, rwa hv at hy end diff --git a/src/topology/algebra/open_subgroup.lean b/src/topology/algebra/open_subgroup.lean index 95b874363d7c0..6f08dd30e8435 100644 --- a/src/topology/algebra/open_subgroup.lean +++ b/src/topology/algebra/open_subgroup.lean @@ -55,44 +55,39 @@ variables {G : Type*} [group G] [topological_space G] variables {U V : open_subgroup G} {g : G} @[to_additive] -instance has_coe_set : has_coe_t (open_subgroup G) (set G) := ⟨λ U, U.1⟩ +instance has_coe_subgroup : has_coe_t (open_subgroup G) (subgroup G) := ⟨to_subgroup⟩ @[to_additive] -instance : has_mem G (open_subgroup G) := ⟨λ g U, g ∈ (U : set G)⟩ +lemma coe_subgroup_injective : injective (coe : open_subgroup G → subgroup G) +| ⟨_, _⟩ ⟨_, _⟩ rfl := rfl @[to_additive] -instance has_coe_subgroup : has_coe_t (open_subgroup G) (subgroup G) := ⟨to_subgroup⟩ +instance : set_like (open_subgroup G) G := +{ coe := λ U, U.1, + coe_injective' := λ _ _ h, coe_subgroup_injective $ set_like.ext' h } + +@[to_additive] +instance : subgroup_class (open_subgroup G) G := +{ mul_mem := λ U _ _, U.mul_mem', + one_mem := λ U, U.one_mem', + inv_mem := λ U _, U.inv_mem' } @[to_additive] instance has_coe_opens : has_coe_t (open_subgroup G) (opens G) := ⟨λ U, ⟨U, U.is_open'⟩⟩ -@[simp, norm_cast, to_additive] lemma mem_coe : g ∈ (U : set G) ↔ g ∈ U := iff.rfl +@[simp, norm_cast, to_additive] lemma coe_coe_opens : ((U : opens G) : set G) = U := rfl +@[simp, norm_cast, to_additive] lemma coe_coe_subgroup : ((U : subgroup G) : set G) = U := rfl @[simp, norm_cast, to_additive] lemma mem_coe_opens : g ∈ (U : opens G) ↔ g ∈ U := iff.rfl @[simp, norm_cast, to_additive] lemma mem_coe_subgroup : g ∈ (U : subgroup G) ↔ g ∈ U := iff.rfl -@[to_additive] lemma coe_injective : injective (coe : open_subgroup G → set G) := -by { rintros ⟨⟨⟩⟩ ⟨⟨⟩⟩ ⟨h⟩, congr, } - @[ext, to_additive] -lemma ext (h : ∀ x, x ∈ U ↔ x ∈ V) : (U = V) := coe_injective $ set.ext h - -@[to_additive] -lemma ext_iff : (U = V) ↔ (∀ x, x ∈ U ↔ x ∈ V) := ⟨λ h x, h ▸ iff.rfl, ext⟩ +lemma ext (h : ∀ x, x ∈ U ↔ x ∈ V) : (U = V) := set_like.ext h variable (U) @[to_additive] protected lemma is_open : is_open (U : set G) := U.is_open' -@[to_additive] -protected lemma one_mem : (1 : G) ∈ U := U.one_mem' - -@[to_additive] -protected lemma inv_mem {g : G} (h : g ∈ U) : g⁻¹ ∈ U := U.inv_mem' h - -@[to_additive] -protected lemma mul_mem {g₁ g₂ : G} (h₁ : g₁ ∈ U) (h₂ : g₂ ∈ U) : g₁ * g₂ ∈ U := U.mul_mem' h₁ h₂ - @[to_additive] lemma mem_nhds_one : (U : set G) ∈ 𝓝 (1 : G) := is_open.mem_nhds U.is_open U.one_mem @@ -101,6 +96,15 @@ variable {U} @[to_additive] instance : has_top (open_subgroup G) := ⟨{ is_open' := is_open_univ, .. (⊤ : subgroup G) }⟩ +@[simp, to_additive] lemma mem_top (x : G) : x ∈ (⊤ : open_subgroup G) := trivial +@[simp, norm_cast, to_additive] lemma coe_top : ((⊤ : open_subgroup G) : set G) = set.univ := rfl + +@[simp, norm_cast, to_additive] +lemma coe_subgroup_top : ((⊤ : open_subgroup G) : subgroup G) = ⊤ := rfl + +@[simp, norm_cast, to_additive] +lemma coe_opens_top : ((⊤ : open_subgroup G) : opens G) = ⊤ := rfl + @[to_additive] instance : inhabited (open_subgroup G) := ⟨⊤⟩ @@ -109,51 +113,60 @@ lemma is_closed [has_continuous_mul G] (U : open_subgroup G) : is_closed (U : se begin apply is_open_compl_iff.1, refine is_open_iff_forall_mem_open.2 (λ x hx, ⟨(λ y, y * x⁻¹) ⁻¹' U, _, _, _⟩), - { intros u hux, - simp only [set.mem_preimage, set.mem_compl_iff, mem_coe] at hux hx ⊢, - refine mt (λ hu, _) hx, + { refine λ u hux hu, hx _, + simp only [set.mem_preimage, set_like.mem_coe] at hux hu ⊢, convert U.mul_mem (U.inv_mem hux) hu, simp }, { exact U.is_open.preimage (continuous_mul_right _) }, - { simp [U.one_mem] } + { simp [one_mem] } end +@[to_additive] +lemma is_clopen [has_continuous_mul G] (U : open_subgroup G) : is_clopen (U : set G) := +⟨U.is_open, U.is_closed⟩ + section variables {H : Type*} [group H] [topological_space H] /-- The product of two open subgroups as an open subgroup of the product group. -/ @[to_additive "The product of two open subgroups as an open subgroup of the product group."] def prod (U : open_subgroup G) (V : open_subgroup H) : open_subgroup (G × H) := -{ carrier := U ×ˢ V, - is_open' := U.is_open.prod V.is_open, +{ is_open' := U.is_open.prod V.is_open, .. (U : subgroup G).prod (V : subgroup H) } +@[simp, norm_cast, to_additive] lemma coe_prod (U : open_subgroup G) (V : open_subgroup H) : + (U.prod V : set (G × H)) = U ×ˢ V := +rfl + +@[simp, norm_cast, to_additive] +lemma coe_subgroup_prod (U : open_subgroup G) (V : open_subgroup H) : + (U.prod V : subgroup (G × H)) = (U : subgroup G).prod V := +rfl + end @[to_additive] -instance : partial_order (open_subgroup G) := -{ le := λ U V, ∀ ⦃x⦄, x ∈ U → x ∈ V, - .. partial_order.lift (coe : open_subgroup G → set G) coe_injective } +instance : has_inf (open_subgroup G) := +⟨λ U V, ⟨U ⊓ V, U.is_open.inter V.is_open⟩⟩ + +@[simp, norm_cast, to_additive] lemma coe_inf : (↑(U ⊓ V) : set G) = (U : set G) ∩ V := rfl +@[simp, norm_cast, to_additive] lemma coe_subgroup_inf : (↑(U ⊓ V) : subgroup G) = ↑U ⊓ ↑V := rfl +@[simp, norm_cast, to_additive] lemma coe_opens_inf : (↑(U ⊓ V) : opens G) = ↑U ⊓ ↑V := rfl +@[simp, to_additive] lemma mem_inf {x} : x ∈ U ⊓ V ↔ x ∈ U ∧ x ∈ V := iff.rfl @[to_additive] instance : semilattice_inf (open_subgroup G) := -{ inf := λ U V, { is_open' := is_open.inter U.is_open V.is_open, .. (U : subgroup G) ⊓ V }, - inf_le_left := λ U V, set.inter_subset_left _ _, - inf_le_right := λ U V, set.inter_subset_right _ _, - le_inf := λ U V W hV hW, set.subset_inter hV hW, - ..open_subgroup.partial_order } +{ .. set_like.partial_order, + .. set_like.coe_injective.semilattice_inf (coe : open_subgroup G → set G) (λ _ _, rfl) } @[to_additive] instance : order_top (open_subgroup G) := { top := ⊤, le_top := λ U, set.subset_univ _ } -@[simp, norm_cast, to_additive] lemma coe_inf : (↑(U ⊓ V) : set G) = (U : set G) ∩ V := rfl - -@[simp, norm_cast, to_additive] lemma coe_subset : (U : set G) ⊆ V ↔ U ≤ V := iff.rfl - @[simp, norm_cast, to_additive] lemma coe_subgroup_le : -(U : subgroup G) ≤ (V : subgroup G) ↔ U ≤ V := iff.rfl + (U : subgroup G) ≤ (V : subgroup G) ↔ U ≤ V := +iff.rfl variables {N : Type*} [group N] [topological_space N] @@ -161,15 +174,18 @@ variables {N : Type*} [group N] [topological_space N] is an `open_subgroup`. -/ @[to_additive "The preimage of an `open_add_subgroup` along a continuous `add_monoid` homomorphism is an `open_add_subgroup`."] -def comap (f : G →* N) - (hf : continuous f) (H : open_subgroup N) : open_subgroup G := +def comap (f : G →* N) (hf : continuous f) (H : open_subgroup N) : open_subgroup G := { is_open' := H.is_open.preimage hf, .. (H : subgroup N).comap f } -@[simp, to_additive] +@[simp, norm_cast, to_additive] lemma coe_comap (H : open_subgroup N) (f : G →* N) (hf : continuous f) : (H.comap f hf : set G) = f ⁻¹' H := rfl +@[simp, norm_cast, to_additive] +lemma coe_subgroup_comap (H : open_subgroup N) (f : G →* N) (hf : continuous f) : + (H.comap f hf : subgroup G) = (H : subgroup N).comap f := rfl + @[simp, to_additive] lemma mem_comap {H : open_subgroup N} {f : G →* N} {hf : continuous f} {x : G} : x ∈ H.comap f hf ↔ f x ∈ H := iff.rfl @@ -190,45 +206,30 @@ variables {G : Type*} [group G] [topological_space G] [has_continuous_mul G] (H lemma is_open_of_mem_nhds {g : G} (hg : (H : set G) ∈ 𝓝 g) : is_open (H : set G) := begin - simp only [is_open_iff_mem_nhds, set_like.mem_coe] at hg ⊢, - intros x hx, - have : filter.tendsto (λ y, y * (x⁻¹ * g)) (𝓝 x) (𝓝 $ x * (x⁻¹ * g)) := - (continuous_id.mul continuous_const).tendsto _, - rw [mul_inv_cancel_left] at this, - have := filter.mem_map'.1 (this hg), - replace hg : g ∈ H := set_like.mem_coe.1 (mem_of_mem_nhds hg), - simp only [set_like.mem_coe, H.mul_mem_cancel_right (H.mul_mem (H.inv_mem hx) hg)] at this, - exact this + refine is_open_iff_mem_nhds.2 (λ x hx, _), + have hg' : g ∈ H := set_like.mem_coe.1 (mem_of_mem_nhds hg), + have : filter.tendsto (λ y, y * (x⁻¹ * g)) (𝓝 x) (𝓝 g) := + (continuous_id.mul continuous_const).tendsto' _ _ (mul_inv_cancel_left _ _), + simpa only [set_like.mem_coe, filter.mem_map', + H.mul_mem_cancel_right (H.mul_mem (H.inv_mem hx) hg')] using this hg, end @[to_additive] -lemma is_open_of_open_subgroup {U : open_subgroup G} (h : U.1 ≤ H) : +lemma is_open_mono {H₁ H₂ : subgroup G} (h : H₁ ≤ H₂) (h₁ : is_open (H₁ : set G)) : + is_open (H₂ : set G) := +is_open_of_mem_nhds _ $ filter.mem_of_superset (h₁.mem_nhds $ one_mem H₁) h + +@[to_additive] +lemma is_open_of_open_subgroup {U : open_subgroup G} (h : ↑U ≤ H) : is_open (H : set G) := -H.is_open_of_mem_nhds (filter.mem_of_superset U.mem_nhds_one h) +is_open_mono h U.is_open /-- If a subgroup of a topological group has `1` in its interior, then it is open. -/ @[to_additive "If a subgroup of an additive topological group has `0` in its interior, then it is open."] -lemma is_open_of_one_mem_interior {G : Type*} [group G] [topological_space G] - [topological_group G] {H : subgroup G} (h_1_int : (1 : G) ∈ interior (H : set G)) : +lemma is_open_of_one_mem_interior (h_1_int : (1 : G) ∈ interior (H : set G)) : is_open (H : set G) := -begin - have h : 𝓝 1 ≤ filter.principal (H : set G) := - nhds_le_of_le h_1_int (is_open_interior) (filter.principal_mono.2 interior_subset), - rw is_open_iff_nhds, - intros g hg, - rw (show 𝓝 g = filter.map ⇑(homeomorph.mul_left g) (𝓝 1), by simp), - convert filter.map_mono h, - simp only [homeomorph.coe_mul_left, filter.map_principal, set.image_mul_left, - filter.principal_eq_iff_eq], - ext, - simp [H.mul_mem_cancel_left (H.inv_mem hg)], -end - -@[to_additive] -lemma is_open_mono {H₁ H₂ : subgroup G} (h : H₁ ≤ H₂) (h₁ : is_open (H₁ : set G)) : - is_open (H₂ : set G) := -@is_open_of_open_subgroup _ _ _ _ H₂ { is_open' := h₁, .. H₁ } h +is_open_of_mem_nhds H $ mem_interior_iff_mem_nhds.1 h_1_int end subgroup @@ -237,20 +238,16 @@ namespace open_subgroup variables {G : Type*} [group G] [topological_space G] [has_continuous_mul G] @[to_additive] -instance : semilattice_sup (open_subgroup G) := -{ sup := λ U V, - { is_open' := show is_open (((U : subgroup G) ⊔ V : subgroup G) : set G), - from subgroup.is_open_mono le_sup_left U.is_open, - .. ((U : subgroup G) ⊔ V) }, - le_sup_left := λ U V, coe_subgroup_le.1 le_sup_left, - le_sup_right := λ U V, coe_subgroup_le.1 le_sup_right, - sup_le := λ U V W hU hV, coe_subgroup_le.1 (sup_le hU hV), - ..open_subgroup.semilattice_inf } +instance : has_sup (open_subgroup G) := +⟨λ U V, ⟨U ⊔ V, subgroup.is_open_mono (le_sup_left : U.1 ≤ U ⊔ V) U.is_open⟩⟩ + +@[simp, norm_cast, to_additive] +lemma coe_subgroup_sup (U V : open_subgroup G) : (↑(U ⊔ V) : subgroup G) = ↑U ⊔ ↑V := rfl @[to_additive] instance : lattice (open_subgroup G) := -{ ..open_subgroup.semilattice_sup, ..open_subgroup.semilattice_inf } - +{ .. open_subgroup.semilattice_inf, + .. coe_subgroup_injective.semilattice_sup (coe : open_subgroup G → subgroup G) (λ _ _, rfl) } end open_subgroup From a37865088599172dc923253bb7b31998297d9c8a Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 15 Mar 2023 14:32:23 +0000 Subject: [PATCH 0633/1291] chore(analysis/inner_product_space/basic): add inner_self_ne_zero (#18587) This result is trivial, but it's presence is consistent with how we have both `norm_ne_zero` and `norm_ne_zero_iff`. --- src/analysis/inner_product_space/basic.lean | 8 +++++++- src/analysis/inner_product_space/dual.lean | 7 +------ src/analysis/inner_product_space/gram_schmidt_ortho.lean | 2 +- src/geometry/euclidean/angle/unoriented/basic.lean | 3 +-- src/geometry/euclidean/basic.lean | 2 +- 5 files changed, 11 insertions(+), 11 deletions(-) diff --git a/src/analysis/inner_product_space/basic.lean b/src/analysis/inner_product_space/basic.lean index bb2dfcdbe93c5..8de1c48ef3df7 100644 --- a/src/analysis/inner_product_space/basic.lean +++ b/src/analysis/inner_product_space/basic.lean @@ -209,6 +209,9 @@ by rw [←inner_conj_symm, inner_zero_left]; simp only [ring_hom.map_zero] lemma inner_self_eq_zero {x : F} : ⟪x, x⟫ = 0 ↔ x = 0 := iff.intro (c.definite _) (by { rintro rfl, exact inner_zero_left _ }) +lemma inner_self_ne_zero {x : F} : ⟪x, x⟫ ≠ 0 ↔ x ≠ 0 := +inner_self_eq_zero.not + lemma inner_self_re_to_K (x : F) : (re ⟪x, x⟫ : 𝕜) = ⟪x, x⟫ := by norm_num [ext_iff, inner_self_nonneg_im] @@ -248,7 +251,7 @@ begin by_cases hy : y = 0, { rw [hy], simp only [is_R_or_C.abs_zero, inner_zero_left, mul_zero, add_monoid_hom.map_zero] }, { change y ≠ 0 at hy, - have hy' : ⟪y, y⟫ ≠ 0 := λ h, by rw [inner_self_eq_zero] at h; exact hy h, + have hy' : ⟪y, y⟫ ≠ 0 := inner_self_ne_zero.mpr hy, set T := ⟪y, x⟫ / ⟪y, y⟫ with hT, have h₁ : re ⟪y, x⟫ = re ⟪x, y⟫ := inner_re_symm _ _, have h₂ : im ⟪y, x⟫ = -im ⟪x, y⟫ := inner_im_symm _ _, @@ -510,6 +513,9 @@ begin exact inner_zero_left _ } end +lemma inner_self_ne_zero {x : E} : ⟪x, x⟫ ≠ 0 ↔ x ≠ 0 := +inner_self_eq_zero.not + @[simp] lemma inner_self_nonpos {x : E} : re ⟪x, x⟫ ≤ 0 ↔ x = 0 := begin split, diff --git a/src/analysis/inner_product_space/dual.lean b/src/analysis/inner_product_space/dual.lean index 16df24f93439b..08d37ec9ac3f1 100644 --- a/src/analysis/inner_product_space/dual.lean +++ b/src/analysis/inner_product_space/dual.lean @@ -131,12 +131,7 @@ begin ... = (ℓ x) * ⟪z, z⟫ / ⟪z, z⟫ : by rw [h₂] ... = ℓ x - : begin - have : ⟪z, z⟫ ≠ 0, - { change z = 0 → false at z_ne_0, - rwa ←inner_self_eq_zero at z_ne_0 }, - field_simp [this] - end, + : by field_simp [inner_self_ne_zero.2 z_ne_0], exact h₄ } end diff --git a/src/analysis/inner_product_space/gram_schmidt_ortho.lean b/src/analysis/inner_product_space/gram_schmidt_ortho.lean index 87cb424373cdb..fbbea0aac7209 100644 --- a/src/analysis/inner_product_space/gram_schmidt_ortho.lean +++ b/src/analysis/inner_product_space/gram_schmidt_ortho.lean @@ -96,7 +96,7 @@ begin { by_cases h : gram_schmidt 𝕜 f a = 0, { simp only [h, inner_zero_left, zero_div, zero_mul, sub_zero], }, { rw [← inner_self_eq_norm_sq_to_K, div_mul_cancel, sub_self], - rwa [ne.def, inner_self_eq_zero], }, }, + rwa [inner_self_ne_zero], }, }, simp_intros i hi hia only [finset.mem_range], simp only [mul_eq_zero, div_eq_zero_iff, inner_self_eq_zero], right, diff --git a/src/geometry/euclidean/angle/unoriented/basic.lean b/src/geometry/euclidean/angle/unoriented/basic.lean index fd4c805d90234..0a3cc92f398bd 100644 --- a/src/geometry/euclidean/angle/unoriented/basic.lean +++ b/src/geometry/euclidean/angle/unoriented/basic.lean @@ -112,8 +112,7 @@ end @[simp] lemma angle_self {x : V} (hx : x ≠ 0) : angle x x = 0 := begin unfold angle, - rw [←real_inner_self_eq_norm_mul_norm, div_self (λ h, hx (inner_self_eq_zero.1 h)), - real.arccos_one] + rw [←real_inner_self_eq_norm_mul_norm, div_self (inner_self_ne_zero.2 hx), real.arccos_one] end /-- The angle between a nonzero vector and its negation. -/ diff --git a/src/geometry/euclidean/basic.lean b/src/geometry/euclidean/basic.lean index e3e284c12be3b..9014fdfc1496b 100644 --- a/src/geometry/euclidean/basic.lean +++ b/src/geometry/euclidean/basic.lean @@ -1046,7 +1046,7 @@ by rw [←neg_one_smul ℝ v, s.second_inter_smul p v (by norm_num : (-1 : ℝ) s.second_inter (s.second_inter p v) v = p := begin by_cases hv : v = 0, { simp [hv] }, - have hv' : ⟪v, v⟫ ≠ 0 := inner_self_eq_zero.not.2 hv, + have hv' : ⟪v, v⟫ ≠ 0 := inner_self_ne_zero.2 hv, simp only [sphere.second_inter, vadd_vsub_assoc, vadd_vadd, inner_add_right, inner_smul_right, div_mul_cancel _ hv'], rw [←@vsub_eq_zero_iff_eq V, vadd_vsub, ←add_smul, ←add_div], From ce7e9d53d4bbc38065db3b595cd5bd73c323bc1d Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 15 Mar 2023 16:02:42 +0000 Subject: [PATCH 0634/1291] feat(algebra/triv_sq_zero_ext): lemmas about big operators (#18488) Some more results following on from #18384. For now this just has the list lemmas. The multiset and finset lemmas are hard to state cleanly. --- src/algebra/triv_sq_zero_ext.lean | 24 ++++++++++++++++++++++++ 1 file changed, 24 insertions(+) diff --git a/src/algebra/triv_sq_zero_ext.lean b/src/algebra/triv_sq_zero_ext.lean index 7cc0bf15019a2..fc2bedc4c9c16 100644 --- a/src/algebra/triv_sq_zero_ext.lean +++ b/src/algebra/triv_sq_zero_ext.lean @@ -494,12 +494,36 @@ instance [monoid R] [add_monoid M] end, .. triv_sq_zero_ext.mul_one_class } +lemma fst_list_prod [monoid R] [add_monoid M] + [distrib_mul_action R M] [distrib_mul_action Rᵐᵒᵖ M] [smul_comm_class R Rᵐᵒᵖ M] + (l : list (tsze R M)) : + l.prod.fst = (l.map fst).prod := +map_list_prod (⟨fst, fst_one, fst_mul⟩ : tsze R M →* R) _ + instance [semiring R] [add_comm_monoid M] [module R M] [module Rᵐᵒᵖ M] [smul_comm_class R Rᵐᵒᵖ M] : semiring (tsze R M) := { .. triv_sq_zero_ext.monoid, .. triv_sq_zero_ext.non_assoc_semiring } +/-- The second element of a product $\prod_{i=0}^n (r_i + m_i)$ is a sum of terms of the form +$r_0\cdots r_{i-1}m_ir_{i+1}\cdots r_n$. -/ +lemma snd_list_prod [semiring R] [add_comm_monoid M] + [module R M] [module Rᵐᵒᵖ M] [smul_comm_class R Rᵐᵒᵖ M] + (l : list (tsze R M)) : + l.prod.snd = + (l.enum.map (λ x : ℕ × tsze R M, + ((l.map fst).take x.1).prod • op ((l.map fst).drop x.1.succ).prod • x.snd.snd)).sum := +begin + induction l with x xs ih, + { simp }, + { rw [list.enum_cons, ←list.map_fst_add_enum_eq_enum_from], + simp_rw [list.map_cons, list.map_map, function.comp, prod.map_snd, prod.map_fst, id, + list.take_zero, list.take_cons, list.prod_nil, list.prod_cons, snd_mul, one_smul, + list.drop, mul_smul, list.sum_cons, fst_list_prod, ih, list.smul_sum, list.map_map], + exact add_comm _ _, } +end + instance [ring R] [add_comm_group M] [module R M] [module Rᵐᵒᵖ M] [smul_comm_class R Rᵐᵒᵖ M] : ring (tsze R M) := From 97d1aa955750bd57a7eeef91de310e633881670b Mon Sep 17 00:00:00 2001 From: Kexing Ying Date: Thu, 16 Mar 2023 06:43:27 +0000 Subject: [PATCH 0635/1291] feat(probability/kernel/invariance): Define pushforward of measure along a kernel (#18244) Co-authored-by: RemyDegenne --- src/measure_theory/measure/measure_space.lean | 5 +- src/probability/kernel/basic.lean | 4 + src/probability/kernel/composition.lean | 2 +- src/probability/kernel/invariance.lean | 155 ++++++++++++++++++ 4 files changed, 164 insertions(+), 2 deletions(-) create mode 100644 src/probability/kernel/invariance.lean diff --git a/src/measure_theory/measure/measure_space.lean b/src/measure_theory/measure/measure_space.lean index 86424d7edcee0..02af7350e7fa0 100644 --- a/src/measure_theory/measure/measure_space.lean +++ b/src/measure_theory/measure/measure_space.lean @@ -695,8 +695,11 @@ instance [measurable_space α] : has_zero (measure α) := @[simp, norm_cast] theorem coe_zero {m : measurable_space α} : ⇑(0 : measure α) = 0 := rfl +instance [is_empty α] {m : measurable_space α} : subsingleton (measure α) := +⟨λ μ ν, by{ ext1 s hs, simp only [eq_empty_of_is_empty s, measure_empty] }⟩ + lemma eq_zero_of_is_empty [is_empty α] {m : measurable_space α} (μ : measure α) : μ = 0 := -ext $ λ s hs, by simp only [eq_empty_of_is_empty s, measure_empty] +subsingleton.elim μ 0 instance [measurable_space α] : inhabited (measure α) := ⟨0⟩ diff --git a/src/probability/kernel/basic.lean b/src/probability/kernel/basic.lean index 22eb09bd1b147..6ebb39aace833 100644 --- a/src/probability/kernel/basic.lean +++ b/src/probability/kernel/basic.lean @@ -357,6 +357,10 @@ def const (α : Type*) {β : Type*} [measurable_space α] {mβ : measurable_spac include mα mβ +lemma const_apply (μβ : measure β) (a : α) : + const α μβ a = μβ := +rfl + instance is_finite_kernel_const {μβ : measure β} [hμβ : is_finite_measure μβ] : is_finite_kernel (const α μβ) := ⟨⟨μβ set.univ, measure_lt_top _ _, λ a, le_rfl⟩⟩ diff --git a/src/probability/kernel/composition.lean b/src/probability/kernel/composition.lean index 68b00d917913f..719367cc72092 100644 --- a/src/probability/kernel/composition.lean +++ b/src/probability/kernel/composition.lean @@ -366,7 +366,7 @@ variables {γ : Type*} {mγ : measurable_space γ} {f : β → γ} {g : γ → include mγ -/-- The pushforward of a kernel along a measurable function. +/-- The pushforward of a kernel along a measurable function. We include measurability in the assumptions instead of using junk values to make sure that typeclass inference can infer that the `map` of a Markov kernel is again a Markov kernel. -/ diff --git a/src/probability/kernel/invariance.lean b/src/probability/kernel/invariance.lean new file mode 100644 index 0000000000000..ddc6ec188d05b --- /dev/null +++ b/src/probability/kernel/invariance.lean @@ -0,0 +1,155 @@ +/- +Copyright (c) 2023 Kexing Ying. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Kexing Ying +-/ +import probability.kernel.composition + +/-! +# Invariance of measures along a kernel + +We define the push-forward of a measure along a kernel which results in another measure. In the +case that the push-forward measure is the same as the original measure, we say that the measure is +invariant with respect to the kernel. + +## Main definitions + +* `probability_theory.kernel.map_measure`: the push-forward of a measure along a kernel. +* `probability_theory.kernel.invariant`: invariance of a given measure with respect to a kernel. + +## Useful lemmas + +* `probability_theory.kernel.comp_apply_eq_map_measure`, + `probability_theory.kernel.const_map_measure_eq_comp_const`, and + `probability_theory.kernel.comp_const_apply_eq_map_measure` established the relationship between + the push-forward measure and the composition of kernels. + +-/ + +open measure_theory + +open_locale measure_theory ennreal probability_theory + +namespace probability_theory + +variables {α β γ : Type*} {mα : measurable_space α} {mβ : measurable_space β} + {mγ : measurable_space γ} + +include mα mβ + +namespace kernel + +/-! ### Push-forward of measures along a kernel -/ + +/-- The push-forward of a measure along a kernel. -/ +noncomputable +def map_measure (κ : kernel α β) (μ : measure α) : + measure β := +measure.of_measurable (λ s hs, ∫⁻ x, κ x s ∂μ) + (by simp only [measure_empty, lintegral_const, zero_mul]) + begin + intros f hf₁ hf₂, + simp_rw [measure_Union hf₂ hf₁, + lintegral_tsum (λ i, (kernel.measurable_coe κ (hf₁ i)).ae_measurable)], + end + +@[simp] +lemma map_measure_apply (κ : kernel α β) (μ : measure α) {s : set β} (hs : measurable_set s) : + map_measure κ μ s = ∫⁻ x, κ x s ∂μ := +by rw [map_measure, measure.of_measurable_apply s hs] + +@[simp] +lemma map_measure_zero (κ : kernel α β) : map_measure κ 0 = 0 := +begin + ext1 s hs, + rw [map_measure_apply κ 0 hs, lintegral_zero_measure, measure.coe_zero, pi.zero_apply], +end + +@[simp] +lemma map_measure_add (κ : kernel α β) (μ ν : measure α) : + map_measure κ (μ + ν) = map_measure κ μ + map_measure κ ν := +begin + ext1 s hs, + rw [map_measure_apply κ (μ + ν) hs, lintegral_add_measure, measure.coe_add, pi.add_apply, + map_measure_apply κ μ hs, map_measure_apply κ ν hs], +end + +@[simp] +lemma map_measure_smul (κ : kernel α β) (μ : measure α) (r : ℝ≥0∞) : + map_measure κ (r • μ) = r • map_measure κ μ := +begin + ext1 s hs, + rw [map_measure_apply κ (r • μ) hs, lintegral_smul_measure, measure.coe_smul, pi.smul_apply, + map_measure_apply κ μ hs, smul_eq_mul], +end + +include mγ + +lemma comp_apply_eq_map_measure + (η : kernel β γ) [is_s_finite_kernel η] (κ : kernel α β) [is_s_finite_kernel κ] (a : α) : + (η ∘ₖ κ) a = map_measure η (κ a) := +begin + ext1 s hs, + rw [comp_apply η κ a hs, map_measure_apply η _ hs], +end + +omit mγ + +lemma const_map_measure_eq_comp_const (κ : kernel α β) [is_s_finite_kernel κ] + (μ : measure α) [is_finite_measure μ] : + const α (map_measure κ μ) = κ ∘ₖ const α μ := +begin + ext1 a, ext1 s hs, + rw [const_apply, map_measure_apply _ _ hs, comp_apply _ _ _ hs, const_apply], +end + +lemma comp_const_apply_eq_map_measure (κ : kernel α β) [is_s_finite_kernel κ] + (μ : measure α) [is_finite_measure μ] (a : α) : + (κ ∘ₖ const α μ) a = map_measure κ μ := +by rw [← const_apply (map_measure κ μ) a, const_map_measure_eq_comp_const κ μ] + +lemma lintegral_map_measure + (κ : kernel α β) [is_s_finite_kernel κ] (μ : measure α) [is_finite_measure μ] + {f : β → ℝ≥0∞} (hf : measurable f) : + ∫⁻ b, f b ∂(map_measure κ μ) = ∫⁻ a, ∫⁻ b, f b ∂(κ a) ∂μ := +begin + by_cases hα : nonempty α, + { have := const_apply μ hα.some, + swap, apply_instance, + conv_rhs { rw [← this] }, + rw [← lintegral_comp _ _ _ hf, ← comp_const_apply_eq_map_measure κ μ hα.some] }, + { haveI := not_nonempty_iff.1 hα, + rw [μ.eq_zero_of_is_empty, map_measure_zero, lintegral_zero_measure, + lintegral_zero_measure] } +end + +omit mβ + +/-! ### Invariant measures of kernels -/ + +/-- A measure `μ` is invariant with respect to the kernel `κ` if the push-forward measure of `μ` +along `κ` equals `μ`. -/ +def invariant (κ : kernel α α) (μ : measure α) : Prop := +map_measure κ μ = μ + +variables {κ η : kernel α α} {μ : measure α} + +lemma invariant.def (hκ : invariant κ μ) : map_measure κ μ = μ := hκ + +lemma invariant.comp_const [is_s_finite_kernel κ] [is_finite_measure μ] + (hκ : invariant κ μ) : (κ ∘ₖ const α μ) = const α μ := +by rw [← const_map_measure_eq_comp_const κ μ, hκ.def] + +lemma invariant.comp [is_s_finite_kernel κ] [is_s_finite_kernel η] [is_finite_measure μ] + (hκ : invariant κ μ) (hη : invariant η μ) : invariant (κ ∘ₖ η) μ := +begin + by_cases hα : nonempty α, + { simp_rw [invariant, ← comp_const_apply_eq_map_measure (κ ∘ₖ η) μ hα.some, comp_assoc, + hη.comp_const, hκ.comp_const, const_apply] }, + { haveI := not_nonempty_iff.1 hα, + exact subsingleton.elim _ _ }, +end + +end kernel + +end probability_theory From 5120cf49cb659e2499edd7e4d336a04efd598f2f Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Thu, 16 Mar 2023 08:10:19 +0000 Subject: [PATCH 0636/1291] chore(data/mv_polynomial): reverse an import (#18593) As proposed on [zulip](https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/!4.232887.20Data.2EMvPolynomial.2EMonad/near/342140597) Co-authored-by: Scott Morrison --- src/data/mv_polynomial/monad.lean | 40 +++++++++++++++++++++++- src/data/mv_polynomial/variables.lean | 38 ++-------------------- src/ring_theory/polynomial/quotient.lean | 19 ++++++----- 3 files changed, 50 insertions(+), 47 deletions(-) diff --git a/src/data/mv_polynomial/monad.lean b/src/data/mv_polynomial/monad.lean index a1b138f35320f..a59fb34f870b7 100644 --- a/src/data/mv_polynomial/monad.lean +++ b/src/data/mv_polynomial/monad.lean @@ -3,8 +3,8 @@ Copyright (c) 2020 Johan Commelin, Robert Y. Lewis. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin, Robert Y. Lewis -/ - import data.mv_polynomial.rename +import data.mv_polynomial.variables /-! @@ -284,6 +284,44 @@ lemma bind₂_monomial_one (f : R →+* mv_polynomial σ S) (d : σ →₀ ℕ) bind₂ f (monomial d 1) = monomial d 1 := by rw [bind₂_monomial, f.map_one, one_mul] +section +open_locale classical + +lemma vars_bind₁ (f : σ → mv_polynomial τ R) (φ : mv_polynomial σ R) : + (bind₁ f φ).vars ⊆ φ.vars.bUnion (λ i, (f i).vars) := +begin + calc (bind₁ f φ).vars + = (φ.support.sum (λ (x : σ →₀ ℕ), (bind₁ f) (monomial x (coeff x φ)))).vars : + by { rw [← alg_hom.map_sum, ← φ.as_sum], } + ... ≤ φ.support.bUnion (λ (i : σ →₀ ℕ), ((bind₁ f) (monomial i (coeff i φ))).vars) : + vars_sum_subset _ _ + ... = φ.support.bUnion (λ (d : σ →₀ ℕ), (C (coeff d φ) * ∏ i in d.support, f i ^ d i).vars) : + by simp only [bind₁_monomial] + ... ≤ φ.support.bUnion (λ (d : σ →₀ ℕ), d.support.bUnion (λ i, (f i).vars)) : _ -- proof below + ... ≤ φ.vars.bUnion (λ (i : σ), (f i).vars) : _, -- proof below + { apply finset.bUnion_mono, + intros d hd, + calc (C (coeff d φ) * ∏ (i : σ) in d.support, f i ^ d i).vars + ≤ (C (coeff d φ)).vars ∪ (∏ (i : σ) in d.support, f i ^ d i).vars : vars_mul _ _ + ... ≤ (∏ (i : σ) in d.support, f i ^ d i).vars : + by simp only [finset.empty_union, vars_C, finset.le_iff_subset, finset.subset.refl] + ... ≤ d.support.bUnion (λ (i : σ), (f i ^ d i).vars) : vars_prod _ + ... ≤ d.support.bUnion (λ (i : σ), (f i).vars) : _, + apply finset.bUnion_mono, + intros i hi, + apply vars_pow, }, + { intro j, + simp_rw finset.mem_bUnion, + rintro ⟨d, hd, ⟨i, hi, hj⟩⟩, + exact ⟨i, (mem_vars _).mpr ⟨d, hd, hi⟩, hj⟩ } +end +end + +lemma mem_vars_bind₁ (f : σ → mv_polynomial τ R) (φ : mv_polynomial σ R) {j : τ} + (h : j ∈ (bind₁ f φ).vars) : + ∃ (i : σ), i ∈ φ.vars ∧ j ∈ (f i).vars := +by simpa only [exists_prop, finset.mem_bUnion, mem_support_iff, ne.def] using vars_bind₁ f φ h + instance monad : monad (λ σ, mv_polynomial σ R) := { map := λ α β f p, rename f p, pure := λ _, X, diff --git a/src/data/mv_polynomial/variables.lean b/src/data/mv_polynomial/variables.lean index b7ce2c7c96b80..7e0447a6f9980 100644 --- a/src/data/mv_polynomial/variables.lean +++ b/src/data/mv_polynomial/variables.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Johan Commelin, Mario Carneiro -/ import algebra.big_operators.order -import data.mv_polynomial.monad +import data.mv_polynomial.rename /-! # Degrees and variables of polynomials @@ -773,7 +773,7 @@ lemma exists_rename_eq_of_vars_subset_range (p : mv_polynomial σ R) (f : τ → σ) (hfi : injective f) (hf : ↑p.vars ⊆ set.range f) : ∃ q : mv_polynomial τ R, rename f q = p := -⟨bind₁ (λ i : σ, option.elim 0 X $ partial_inv f i) p, +⟨aeval (λ i : σ, option.elim 0 X $ partial_inv f i) p, begin show (rename f).to_ring_hom.comp _ p = ring_hom.id _ p, refine hom_congr_vars _ _ _, @@ -785,40 +785,6 @@ lemma exists_rename_eq_of_vars_subset_range { refl } end⟩ -lemma vars_bind₁ (f : σ → mv_polynomial τ R) (φ : mv_polynomial σ R) : - (bind₁ f φ).vars ⊆ φ.vars.bUnion (λ i, (f i).vars) := -begin - calc (bind₁ f φ).vars - = (φ.support.sum (λ (x : σ →₀ ℕ), (bind₁ f) (monomial x (coeff x φ)))).vars : - by { rw [← alg_hom.map_sum, ← φ.as_sum], } - ... ≤ φ.support.bUnion (λ (i : σ →₀ ℕ), ((bind₁ f) (monomial i (coeff i φ))).vars) : - vars_sum_subset _ _ - ... = φ.support.bUnion (λ (d : σ →₀ ℕ), (C (coeff d φ) * ∏ i in d.support, f i ^ d i).vars) : - by simp only [bind₁_monomial] - ... ≤ φ.support.bUnion (λ (d : σ →₀ ℕ), d.support.bUnion (λ i, (f i).vars)) : _ -- proof below - ... ≤ φ.vars.bUnion (λ (i : σ), (f i).vars) : _, -- proof below - { apply finset.bUnion_mono, - intros d hd, - calc (C (coeff d φ) * ∏ (i : σ) in d.support, f i ^ d i).vars - ≤ (C (coeff d φ)).vars ∪ (∏ (i : σ) in d.support, f i ^ d i).vars : vars_mul _ _ - ... ≤ (∏ (i : σ) in d.support, f i ^ d i).vars : - by simp only [finset.empty_union, vars_C, finset.le_iff_subset, finset.subset.refl] - ... ≤ d.support.bUnion (λ (i : σ), (f i ^ d i).vars) : vars_prod _ - ... ≤ d.support.bUnion (λ (i : σ), (f i).vars) : _, - apply finset.bUnion_mono, - intros i hi, - apply vars_pow, }, - { intro j, - simp_rw finset.mem_bUnion, - rintro ⟨d, hd, ⟨i, hi, hj⟩⟩, - exact ⟨i, (mem_vars _).mpr ⟨d, hd, hi⟩, hj⟩ } -end - -lemma mem_vars_bind₁ (f : σ → mv_polynomial τ R) (φ : mv_polynomial σ R) {j : τ} - (h : j ∈ (bind₁ f φ).vars) : - ∃ (i : σ), i ∈ φ.vars ∧ j ∈ (f i).vars := -by simpa only [exists_prop, finset.mem_bUnion, mem_support_iff, ne.def] using vars_bind₁ f φ h - lemma vars_rename (f : σ → τ) (φ : mv_polynomial σ R) : (rename f φ).vars ⊆ (φ.vars.image f) := begin diff --git a/src/ring_theory/polynomial/quotient.lean b/src/ring_theory/polynomial/quotient.lean index 155a1a2c368c8..ba8f0338e9692 100644 --- a/src/ring_theory/polynomial/quotient.lean +++ b/src/ring_theory/polynomial/quotient.lean @@ -201,14 +201,14 @@ def quotient_equiv_quotient_mv_polynomial (I : ideal R) : apply induction_on f, { rintro ⟨r⟩, rw [coe_eval₂_hom, eval₂_C], - simp only [eval₂_hom_eq_bind₂, submodule.quotient.quot_mk_eq_mk, ideal.quotient.lift_mk, - ideal.quotient.mk_eq_mk, bind₂_C_right, ring_hom.coe_comp] }, + simp only [submodule.quotient.quot_mk_eq_mk, ideal.quotient.lift_mk, + mv_polynomial.eval₂_hom_C, function.comp_app, ideal.quotient.mk_eq_mk, mv_polynomial.C_inj, + ring_hom.coe_comp], }, { simp_intros p q hp hq only [ring_hom.map_add, mv_polynomial.coe_eval₂_hom, coe_eval₂_hom, - mv_polynomial.eval₂_add, mv_polynomial.eval₂_hom_eq_bind₂, eval₂_hom_eq_bind₂], + mv_polynomial.eval₂_add], rw [hp, hq] }, - { simp_intros p i hp only [eval₂_hom_eq_bind₂, coe_eval₂_hom], - simp only [hp, eval₂_hom_eq_bind₂, coe_eval₂_hom, ideal.quotient.lift_mk, bind₂_X_right, - eval₂_mul, ring_hom.map_mul, eval₂_X] } + { simp_intros p i hp only [coe_eval₂_hom], + simp only [hp, coe_eval₂_hom, ideal.quotient.lift_mk, eval₂_mul, ring_hom.map_mul, eval₂_X] } end, right_inv := begin rintro ⟨f⟩, @@ -216,12 +216,11 @@ def quotient_equiv_quotient_mv_polynomial (I : ideal R) : { intros r, simp only [submodule.quotient.quot_mk_eq_mk, ideal.quotient.lift_mk, ideal.quotient.mk_eq_mk, ring_hom.coe_comp, eval₂_hom_C] }, - { simp_intros p q hp hq only [eval₂_hom_eq_bind₂, submodule.quotient.quot_mk_eq_mk, eval₂_add, + { simp_intros p q hp hq only [submodule.quotient.quot_mk_eq_mk, eval₂_add, ring_hom.map_add, coe_eval₂_hom, ideal.quotient.lift_mk, ideal.quotient.mk_eq_mk], rw [hp, hq] }, - { simp_intros p i hp only [eval₂_hom_eq_bind₂, submodule.quotient.quot_mk_eq_mk, coe_eval₂_hom, - ideal.quotient.lift_mk, ideal.quotient.mk_eq_mk, bind₂_X_right, eval₂_mul, ring_hom.map_mul, - eval₂_X], + { simp_intros p i hp only [submodule.quotient.quot_mk_eq_mk, coe_eval₂_hom, + ideal.quotient.lift_mk, ideal.quotient.mk_eq_mk, eval₂_mul, ring_hom.map_mul, eval₂_X], simp only [hp] } end, commutes' := λ r, eval₂_hom_C _ _ (ideal.quotient.mk I r) } From eea141bc9cf205beebfd46e2068c7c01ee8db4f6 Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Thu, 16 Mar 2023 11:15:19 +0000 Subject: [PATCH 0637/1291] refactor(geometry/euclidean): split out spheres (#18220) Rearrange material related to spheres as follows: * `geometry.euclidean.sphere.power` has most of the content from the previous `geometry.euclidean.sphere` (intersecting chords / secants, which at some point should probably be refactored using an explicit definition of the power of a point). * `geometry.euclidean.sphere.ptolemy` has the proof of Ptolemy's theorem that was previously in `geometry.euclidean.sphere`. * `geometry.euclidean.sphere.basic` has most of the material previously in `geometry.euclidean.basic`: definitions of `sphere`, `cospherical` and `concyclic` and associated lemmas. * `geometry.euclidean.sphere.second_inter` has the definition and lemmas about `second_inter`. There are no changes to API or proofs. --- archive/imo/imo2019_q2.lean | 1 + src/geometry/euclidean/basic.lean | 464 ------------------ src/geometry/euclidean/circumcenter.lean | 2 +- src/geometry/euclidean/sphere/basic.lean | 349 +++++++++++++ .../{sphere.lean => sphere/power.lean} | 54 +- src/geometry/euclidean/sphere/ptolemy.lean | 72 +++ .../euclidean/sphere/second_inter.lean | 175 +++++++ 7 files changed, 603 insertions(+), 514 deletions(-) create mode 100644 src/geometry/euclidean/sphere/basic.lean rename src/geometry/euclidean/{sphere.lean => sphere/power.lean} (64%) create mode 100644 src/geometry/euclidean/sphere/ptolemy.lean create mode 100644 src/geometry/euclidean/sphere/second_inter.lean diff --git a/archive/imo/imo2019_q2.lean b/archive/imo/imo2019_q2.lean index de839c4451984..d2af904f035d4 100644 --- a/archive/imo/imo2019_q2.lean +++ b/archive/imo/imo2019_q2.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Joseph Myers -/ import geometry.euclidean.angle.sphere +import geometry.euclidean.sphere.second_inter /-! # IMO 2019 Q2 diff --git a/src/geometry/euclidean/basic.lean b/src/geometry/euclidean/basic.lean index 9014fdfc1496b..f2e3bb48c6317 100644 --- a/src/geometry/euclidean/basic.lean +++ b/src/geometry/euclidean/basic.lean @@ -3,7 +3,6 @@ Copyright (c) 2020 Joseph Myers. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Joseph Myers, Manuel Candales -/ -import analysis.convex.strict_convex_between import analysis.inner_product_space.projection import algebra.quadratic_discriminant @@ -665,467 +664,4 @@ lemma reflection_vadd_smul_vsub_orthogonal_projection {s : affine_subspace ℝ P reflection_orthogonal_vadd hp₁ (submodule.smul_mem _ _ (vsub_orthogonal_projection_mem_direction_orthogonal s _)) -omit V - -variables (P) - -/-- A `sphere P` bundles a `center` and `radius`. This definition does not require the radius to -be positive; that should be given as a hypothesis to lemmas that require it. -/ -@[ext] structure sphere := -(center : P) -(radius : ℝ) - -variables {P} - -instance [nonempty P] : nonempty (sphere P) := ⟨⟨classical.arbitrary P, 0⟩⟩ - -instance : has_coe (sphere P) (set P) := ⟨λ s, metric.sphere s.center s.radius⟩ -instance : has_mem P (sphere P) := ⟨λ p s, p ∈ (s : set P)⟩ - -lemma sphere.mk_center (c : P) (r : ℝ) : (⟨c, r⟩ : sphere P).center = c := rfl - -lemma sphere.mk_radius (c : P) (r : ℝ) : (⟨c, r⟩ : sphere P).radius = r := rfl - -@[simp] lemma sphere.mk_center_radius (s : sphere P) : (⟨s.center, s.radius⟩ : sphere P) = s := -by ext; refl - -lemma sphere.coe_def (s : sphere P) : (s : set P) = metric.sphere s.center s.radius := rfl - -@[simp] lemma sphere.coe_mk (c : P) (r : ℝ) : ↑(⟨c, r⟩ : sphere P) = metric.sphere c r := rfl - -@[simp] lemma sphere.mem_coe {p : P} {s : sphere P} : p ∈ (s : set P) ↔ p ∈ s := iff.rfl - -lemma mem_sphere {p : P} {s : sphere P} : p ∈ s ↔ dist p s.center = s.radius := iff.rfl - -lemma mem_sphere' {p : P} {s : sphere P} : p ∈ s ↔ dist s.center p = s.radius := -metric.mem_sphere' - -lemma subset_sphere {ps : set P} {s : sphere P} : ps ⊆ s ↔ ∀ p ∈ ps, p ∈ s := iff.rfl - -lemma dist_of_mem_subset_sphere {p : P} {ps : set P} {s : sphere P} (hp : p ∈ ps) - (hps : ps ⊆ (s : set P)) : dist p s.center = s.radius := -mem_sphere.1 (sphere.mem_coe.1 (set.mem_of_mem_of_subset hp hps)) - -lemma dist_of_mem_subset_mk_sphere {p c : P} {ps : set P} {r : ℝ} (hp : p ∈ ps) - (hps : ps ⊆ ↑(⟨c, r⟩ : sphere P)) : dist p c = r := -dist_of_mem_subset_sphere hp hps - -lemma sphere.ne_iff {s₁ s₂ : sphere P} : - s₁ ≠ s₂ ↔ s₁.center ≠ s₂.center ∨ s₁.radius ≠ s₂.radius := -by rw [←not_and_distrib, ←sphere.ext_iff] - -lemma sphere.center_eq_iff_eq_of_mem {s₁ s₂ : sphere P} {p : P} (hs₁ : p ∈ s₁) (hs₂ : p ∈ s₂) : - s₁.center = s₂.center ↔ s₁ = s₂ := -begin - refine ⟨λ h, sphere.ext _ _ h _, λ h, h ▸ rfl⟩, - rw mem_sphere at hs₁ hs₂, - rw [←hs₁, ←hs₂, h] -end - -lemma sphere.center_ne_iff_ne_of_mem {s₁ s₂ : sphere P} {p : P} (hs₁ : p ∈ s₁) (hs₂ : p ∈ s₂) : - s₁.center ≠ s₂.center ↔ s₁ ≠ s₂ := -(sphere.center_eq_iff_eq_of_mem hs₁ hs₂).not - -lemma dist_center_eq_dist_center_of_mem_sphere {p₁ p₂ : P} {s : sphere P} (hp₁ : p₁ ∈ s) - (hp₂ : p₂ ∈ s) : dist p₁ s.center = dist p₂ s.center := -by rw [mem_sphere.1 hp₁, mem_sphere.1 hp₂] - -lemma dist_center_eq_dist_center_of_mem_sphere' {p₁ p₂ : P} {s : sphere P} (hp₁ : p₁ ∈ s) - (hp₂ : p₂ ∈ s) : dist s.center p₁ = dist s.center p₂ := -by rw [mem_sphere'.1 hp₁, mem_sphere'.1 hp₂] - -/-- A set of points is cospherical if they are equidistant from some -point. In two dimensions, this is the same thing as being -concyclic. -/ -def cospherical (ps : set P) : Prop := -∃ (center : P) (radius : ℝ), ∀ p ∈ ps, dist p center = radius - -/-- The definition of `cospherical`. -/ -lemma cospherical_def (ps : set P) : - cospherical ps ↔ ∃ (center : P) (radius : ℝ), ∀ p ∈ ps, dist p center = radius := -iff.rfl - -/-- A set of points is cospherical if and only if they lie in some sphere. -/ -lemma cospherical_iff_exists_sphere {ps : set P} : - cospherical ps ↔ ∃ s : sphere P, ps ⊆ (s : set P) := -begin - refine ⟨λ h, _, λ h, _⟩, - { rcases h with ⟨c, r, h⟩, - exact ⟨⟨c, r⟩, h⟩ }, - { rcases h with ⟨s, h⟩, - exact ⟨s.center, s.radius, h⟩ } -end - -/-- The set of points in a sphere is cospherical. -/ -lemma sphere.cospherical (s : sphere P) : cospherical (s : set P) := -cospherical_iff_exists_sphere.2 ⟨s, set.subset.rfl⟩ - -/-- A subset of a cospherical set is cospherical. -/ -lemma cospherical.subset {ps₁ ps₂ : set P} (hs : ps₁ ⊆ ps₂) (hc : cospherical ps₂) : - cospherical ps₁ := -begin - rcases hc with ⟨c, r, hcr⟩, - exact ⟨c, r, λ p hp, hcr p (hs hp)⟩ -end - -include V - -/-- The empty set is cospherical. -/ -lemma cospherical_empty : cospherical (∅ : set P) := -begin - use add_torsor.nonempty.some, - simp, -end - -omit V - -/-- A single point is cospherical. -/ -lemma cospherical_singleton (p : P) : cospherical ({p} : set P) := -begin - use p, - simp -end - -include V - -/-- Two points are cospherical. -/ -lemma cospherical_pair (p₁ p₂ : P) : cospherical ({p₁, p₂} : set P) := -begin - use [(2⁻¹ : ℝ) • (p₂ -ᵥ p₁) +ᵥ p₁, (2⁻¹ : ℝ) * (dist p₂ p₁)], - intro p, - rw [set.mem_insert_iff, set.mem_singleton_iff], - rintro ⟨_|_⟩, - { rw [dist_eq_norm_vsub V p₁, vsub_vadd_eq_vsub_sub, vsub_self, zero_sub, norm_neg, norm_smul, - dist_eq_norm_vsub V p₂], - simp }, - { rw [H, dist_eq_norm_vsub V p₂, vsub_vadd_eq_vsub_sub, dist_eq_norm_vsub V p₂], - conv_lhs { congr, congr, rw ←one_smul ℝ (p₂ -ᵥ p₁ : V) }, - rw [←sub_smul, norm_smul], - norm_num } -end - -/-- Any three points in a cospherical set are affinely independent. -/ -lemma cospherical.affine_independent {s : set P} (hs : cospherical s) {p : fin 3 → P} - (hps : set.range p ⊆ s) (hpi : function.injective p) : - affine_independent ℝ p := -begin - rw affine_independent_iff_not_collinear, - intro hc, - rw collinear_iff_of_mem (set.mem_range_self (0 : fin 3)) at hc, - rcases hc with ⟨v, hv⟩, - rw set.forall_range_iff at hv, - have hv0 : v ≠ 0, - { intro h, - have he : p 1 = p 0, by simpa [h] using hv 1, - exact (dec_trivial : (1 : fin 3) ≠ 0) (hpi he) }, - rcases hs with ⟨c, r, hs⟩, - have hs' := λ i, hs (p i) (set.mem_of_mem_of_subset (set.mem_range_self _) hps), - choose f hf using hv, - have hsd : ∀ i, dist ((f i • v) +ᵥ p 0) c = r, - { intro i, - rw ←hf, - exact hs' i }, - have hf0 : f 0 = 0, - { have hf0' := hf 0, - rw [eq_comm, ←@vsub_eq_zero_iff_eq V, vadd_vsub, smul_eq_zero] at hf0', - simpa [hv0] using hf0' }, - have hfi : function.injective f, - { intros i j h, - have hi := hf i, - rw [h, ←hf j] at hi, - exact hpi hi }, - simp_rw [←hsd 0, hf0, zero_smul, zero_vadd, dist_smul_vadd_eq_dist (p 0) c hv0] at hsd, - have hfn0 : ∀ i, i ≠ 0 → f i ≠ 0 := λ i, (hfi.ne_iff' hf0).2, - have hfn0' : ∀ i, i ≠ 0 → f i = (-2) * ⟪v, (p 0 -ᵥ c)⟫ / ⟪v, v⟫, - { intros i hi, - have hsdi := hsd i, - simpa [hfn0, hi] using hsdi }, - have hf12 : f 1 = f 2, { rw [hfn0' 1 dec_trivial, hfn0' 2 dec_trivial] }, - exact (dec_trivial : (1 : fin 3) ≠ 2) (hfi hf12) -end - -/-- Any three points in a cospherical set are affinely independent. -/ -lemma cospherical.affine_independent_of_mem_of_ne {s : set P} (hs : cospherical s) {p₁ p₂ p₃ : P} - (h₁ : p₁ ∈ s) (h₂ : p₂ ∈ s) (h₃ : p₃ ∈ s) (h₁₂ : p₁ ≠ p₂) (h₁₃ : p₁ ≠ p₃) (h₂₃ : p₂ ≠ p₃) : - affine_independent ℝ ![p₁, p₂, p₃] := -begin - refine hs.affine_independent _ _, - { simp [h₁, h₂, h₃, set.insert_subset] }, - { erw [fin.cons_injective_iff, fin.cons_injective_iff], - simp [h₁₂, h₁₃, h₂₃, function.injective] } -end - -/-- The three points of a cospherical set are affinely independent. -/ -lemma cospherical.affine_independent_of_ne {p₁ p₂ p₃ : P} (hs : cospherical ({p₁, p₂, p₃} : set P)) - (h₁₂ : p₁ ≠ p₂) (h₁₃ : p₁ ≠ p₃) (h₂₃ : p₂ ≠ p₃) : - affine_independent ℝ ![p₁, p₂, p₃] := -hs.affine_independent_of_mem_of_ne (set.mem_insert _ _) - (set.mem_insert_of_mem _ (set.mem_insert _ _)) - (set.mem_insert_of_mem _ (set.mem_insert_of_mem _ (set.mem_singleton _))) h₁₂ h₁₃ h₂₃ - -/-- Suppose that `p₁` and `p₂` lie in spheres `s₁` and `s₂`. Then the vector between the centers -of those spheres is orthogonal to that between `p₁` and `p₂`; this is a version of -`inner_vsub_vsub_of_dist_eq_of_dist_eq` for bundled spheres. (In two dimensions, this says that -the diagonals of a kite are orthogonal.) -/ -lemma inner_vsub_vsub_of_mem_sphere_of_mem_sphere {p₁ p₂ : P} {s₁ s₂ : sphere P} - (hp₁s₁ : p₁ ∈ s₁) (hp₂s₁ : p₂ ∈ s₁) (hp₁s₂ : p₁ ∈ s₂) (hp₂s₂ : p₂ ∈ s₂) : - ⟪s₂.center -ᵥ s₁.center, p₂ -ᵥ p₁⟫ = 0 := -inner_vsub_vsub_of_dist_eq_of_dist_eq (dist_center_eq_dist_center_of_mem_sphere hp₁s₁ hp₂s₁) - (dist_center_eq_dist_center_of_mem_sphere hp₁s₂ hp₂s₂) - -/-- Two spheres intersect in at most two points in a two-dimensional subspace containing their -centers; this is a version of `eq_of_dist_eq_of_dist_eq_of_mem_of_finrank_eq_two` for bundled -spheres. -/ -lemma eq_of_mem_sphere_of_mem_sphere_of_mem_of_finrank_eq_two {s : affine_subspace ℝ P} - [finite_dimensional ℝ s.direction] (hd : finrank ℝ s.direction = 2) {s₁ s₂ : sphere P} - {p₁ p₂ p : P} (hs₁ : s₁.center ∈ s) (hs₂ : s₂.center ∈ s) (hp₁s : p₁ ∈ s) (hp₂s : p₂ ∈ s) - (hps : p ∈ s) (hs : s₁ ≠ s₂) (hp : p₁ ≠ p₂) (hp₁s₁ : p₁ ∈ s₁) (hp₂s₁ : p₂ ∈ s₁) (hps₁ : p ∈ s₁) - (hp₁s₂ : p₁ ∈ s₂) (hp₂s₂ : p₂ ∈ s₂) (hps₂ : p ∈ s₂) : p = p₁ ∨ p = p₂ := -eq_of_dist_eq_of_dist_eq_of_mem_of_finrank_eq_two hd hs₁ hs₂ hp₁s hp₂s hps - ((sphere.center_ne_iff_ne_of_mem hps₁ hps₂).2 hs) hp hp₁s₁ hp₂s₁ hps₁ hp₁s₂ hp₂s₂ hps₂ - -/-- Two spheres intersect in at most two points in two-dimensional space; this is a version of -`eq_of_dist_eq_of_dist_eq_of_finrank_eq_two` for bundled spheres. -/ -lemma eq_of_mem_sphere_of_mem_sphere_of_finrank_eq_two [finite_dimensional ℝ V] - (hd : finrank ℝ V = 2) {s₁ s₂ : sphere P} {p₁ p₂ p : P} (hs : s₁ ≠ s₂) (hp : p₁ ≠ p₂) - (hp₁s₁ : p₁ ∈ s₁) (hp₂s₁ : p₂ ∈ s₁) (hps₁ : p ∈ s₁) (hp₁s₂ : p₁ ∈ s₂) (hp₂s₂ : p₂ ∈ s₂) - (hps₂ : p ∈ s₂) : p = p₁ ∨ p = p₂ := -eq_of_dist_eq_of_dist_eq_of_finrank_eq_two hd ((sphere.center_ne_iff_ne_of_mem hps₁ hps₂).2 hs) - hp hp₁s₁ hp₂s₁ hps₁ hp₁s₂ hp₂s₂ hps₂ - -/-- Given a point on a sphere and a point not outside it, the inner product between the -difference of those points and the radius vector is positive unless the points are equal. -/ -lemma inner_pos_or_eq_of_dist_le_radius {s : sphere P} {p₁ p₂ : P} (hp₁ : p₁ ∈ s) - (hp₂ : dist p₂ s.center ≤ s.radius) : 0 < ⟪p₁ -ᵥ p₂, p₁ -ᵥ s.center⟫ ∨ p₁ = p₂ := -begin - by_cases h : p₁ = p₂, { exact or.inr h }, - refine or.inl _, - rw mem_sphere at hp₁, - rw [←vsub_sub_vsub_cancel_right p₁ p₂ s.center, inner_sub_left, - real_inner_self_eq_norm_mul_norm/-, ←dist_eq_norm_vsub, hp₁-/, sub_pos], - refine lt_of_le_of_ne - ((real_inner_le_norm _ _).trans (mul_le_mul_of_nonneg_right _ (norm_nonneg _))) - _, - { rwa [←dist_eq_norm_vsub, ←dist_eq_norm_vsub, hp₁] }, - { rcases hp₂.lt_or_eq with hp₂' | hp₂', - { refine ((real_inner_le_norm _ _).trans_lt (mul_lt_mul_of_pos_right _ _)).ne, - { rwa [←hp₁, @dist_eq_norm_vsub V, @dist_eq_norm_vsub V] at hp₂' }, - { rw [norm_pos_iff, vsub_ne_zero], - rintro rfl, - rw ←hp₁ at hp₂', - refine (dist_nonneg.not_lt : ¬dist p₂ s.center < 0) _, - simpa using hp₂' } }, - { rw [←hp₁, @dist_eq_norm_vsub V, @dist_eq_norm_vsub V] at hp₂', - nth_rewrite 0 ←hp₂', - rw [ne.def, inner_eq_norm_mul_iff_real, hp₂', ←sub_eq_zero, ←smul_sub, - vsub_sub_vsub_cancel_right, ←ne.def, smul_ne_zero_iff, vsub_ne_zero, - and_iff_left (ne.symm h), norm_ne_zero_iff, vsub_ne_zero], - rintro rfl, - refine h (eq.symm _), - simpa using hp₂' } } -end - -/-- Given a point on a sphere and a point not outside it, the inner product between the -difference of those points and the radius vector is nonnegative. -/ -lemma inner_nonneg_of_dist_le_radius {s : sphere P} {p₁ p₂ : P} (hp₁ : p₁ ∈ s) - (hp₂ : dist p₂ s.center ≤ s.radius) : 0 ≤ ⟪p₁ -ᵥ p₂, p₁ -ᵥ s.center⟫ := -begin - rcases inner_pos_or_eq_of_dist_le_radius hp₁ hp₂ with h | rfl, - { exact h.le }, - { simp } -end - -/-- Given a point on a sphere and a point inside it, the inner product between the difference of -those points and the radius vector is positive. -/ -lemma inner_pos_of_dist_lt_radius {s : sphere P} {p₁ p₂ : P} (hp₁ : p₁ ∈ s) - (hp₂ : dist p₂ s.center < s.radius) : 0 < ⟪p₁ -ᵥ p₂, p₁ -ᵥ s.center⟫ := -begin - by_cases h : p₁ = p₂, - { rw [h, mem_sphere] at hp₁, - exact false.elim (hp₂.ne hp₁) }, - exact (inner_pos_or_eq_of_dist_le_radius hp₁ hp₂.le).resolve_right h -end - -/-- Given three collinear points, two on a sphere and one not outside it, the one not outside it -is weakly between the other two points. -/ -lemma wbtw_of_collinear_of_dist_center_le_radius {s : sphere P} {p₁ p₂ p₃ : P} - (h : collinear ℝ ({p₁, p₂, p₃} : set P)) (hp₁ : p₁ ∈ s) (hp₂ : dist p₂ s.center ≤ s.radius) - (hp₃ : p₃ ∈ s) (hp₁p₃ : p₁ ≠ p₃) : wbtw ℝ p₁ p₂ p₃ := -h.wbtw_of_dist_eq_of_dist_le hp₁ hp₂ hp₃ hp₁p₃ - -/-- Given three collinear points, two on a sphere and one inside it, the one inside it is -strictly between the other two points. -/ -lemma sbtw_of_collinear_of_dist_center_lt_radius {s : sphere P} {p₁ p₂ p₃ : P} - (h : collinear ℝ ({p₁, p₂, p₃} : set P)) (hp₁ : p₁ ∈ s) (hp₂ : dist p₂ s.center < s.radius) - (hp₃ : p₃ ∈ s) (hp₁p₃ : p₁ ≠ p₃) : sbtw ℝ p₁ p₂ p₃ := -h.sbtw_of_dist_eq_of_dist_lt hp₁ hp₂ hp₃ hp₁p₃ - -/-- The second intersection of a sphere with a line through a point on that sphere; that point -if it is the only point of intersection of the line with the sphere. The intended use of this -definition is when `p ∈ s`; the definition does not use `s.radius`, so in general it returns -the second intersection with the sphere through `p` and with center `s.center`. -/ -def sphere.second_inter (s : sphere P) (p : P) (v : V) : P := -(-2 * ⟪v, p -ᵥ s.center⟫ / ⟪v, v⟫) • v +ᵥ p - -/-- The distance between `second_inter` and the center equals the distance between the original -point and the center. -/ -@[simp] lemma sphere.second_inter_dist (s : sphere P) (p : P) (v : V) : - dist (s.second_inter p v) s.center = dist p s.center := -begin - rw sphere.second_inter, - by_cases hv : v = 0, { simp [hv] }, - rw dist_smul_vadd_eq_dist _ _ hv, - exact or.inr rfl -end - -/-- The point given by `second_inter` lies on the sphere. -/ -@[simp] lemma sphere.second_inter_mem {s : sphere P} {p : P} (v : V) : - s.second_inter p v ∈ s ↔ p ∈ s := -by simp_rw [mem_sphere, sphere.second_inter_dist] - -variables (V) - -/-- If the vector is zero, `second_inter` gives the original point. -/ -@[simp] lemma sphere.second_inter_zero (s : sphere P) (p : P) : - s.second_inter p (0 : V) = p := -by simp [sphere.second_inter] - -variables {V} - -/-- The point given by `second_inter` equals the original point if and only if the line is -orthogonal to the radius vector. -/ -lemma sphere.second_inter_eq_self_iff {s : sphere P} {p : P} {v : V} : - s.second_inter p v = p ↔ ⟪v, p -ᵥ s.center⟫ = 0 := -begin - refine ⟨λ hp, _, λ hp, _⟩, - { by_cases hv : v = 0, { simp [hv] }, - rwa [sphere.second_inter, eq_comm, eq_vadd_iff_vsub_eq, vsub_self, eq_comm, smul_eq_zero, - or_iff_left hv, div_eq_zero_iff, inner_self_eq_zero, or_iff_left hv, mul_eq_zero, - or_iff_right (by norm_num : (-2 : ℝ) ≠ 0)] at hp }, - { rw [sphere.second_inter, hp, mul_zero, zero_div, zero_smul, zero_vadd] } -end - -/-- A point on a line through a point on a sphere equals that point or `second_inter`. -/ -lemma sphere.eq_or_eq_second_inter_of_mem_mk'_span_singleton_iff_mem {s : sphere P} {p : P} - (hp : p ∈ s) {v : V} {p' : P} (hp' : p' ∈ affine_subspace.mk' p (ℝ ∙ v)) : - (p' = p ∨ p' = s.second_inter p v) ↔ p' ∈ s := -begin - refine ⟨λ h, _, λ h, _⟩, - { rcases h with h | h, - { rwa h }, - { rwa [h, sphere.second_inter_mem] } }, - { rw [affine_subspace.mem_mk'_iff_vsub_mem, submodule.mem_span_singleton] at hp', - rcases hp' with ⟨r, hr⟩, - rw [eq_comm, ←eq_vadd_iff_vsub_eq] at hr, - subst hr, - by_cases hv : v = 0, { simp [hv] }, - rw sphere.second_inter, - rw mem_sphere at h hp, - rw [←hp, dist_smul_vadd_eq_dist _ _ hv] at h, - rcases h with h | h; - simp [h] } -end - -/-- `second_inter` is unchanged by multiplying the vector by a nonzero real. -/ -@[simp] lemma sphere.second_inter_smul (s : sphere P) (p : P) (v : V) {r : ℝ} - (hr : r ≠ 0) : s.second_inter p (r • v) = s.second_inter p v := -begin - simp_rw [sphere.second_inter, real_inner_smul_left, inner_smul_right, smul_smul, - div_mul_eq_div_div], - rw [mul_comm, ←mul_div_assoc, ←mul_div_assoc, mul_div_cancel_left _ hr, mul_comm, mul_assoc, - mul_div_cancel_left _ hr, mul_comm] -end - -/-- `second_inter` is unchanged by negating the vector. -/ -@[simp] lemma sphere.second_inter_neg (s : sphere P) (p : P) (v : V) : - s.second_inter p (-v) = s.second_inter p v := -by rw [←neg_one_smul ℝ v, s.second_inter_smul p v (by norm_num : (-1 : ℝ) ≠ 0)] - -/-- Applying `second_inter` twice returns the original point. -/ -@[simp] lemma sphere.second_inter_second_inter (s : sphere P) (p : P) (v : V) : - s.second_inter (s.second_inter p v) v = p := -begin - by_cases hv : v = 0, { simp [hv] }, - have hv' : ⟪v, v⟫ ≠ 0 := inner_self_ne_zero.2 hv, - simp only [sphere.second_inter, vadd_vsub_assoc, vadd_vadd, inner_add_right, inner_smul_right, - div_mul_cancel _ hv'], - rw [←@vsub_eq_zero_iff_eq V, vadd_vsub, ←add_smul, ←add_div], - convert zero_smul ℝ _, - convert zero_div _, - ring -end - -/-- If the vector passed to `second_inter` is given by a subtraction involving the point in -`second_inter`, the result of `second_inter` may be expressed using `line_map`. -/ -lemma sphere.second_inter_eq_line_map (s : sphere P) (p p' : P) : - s.second_inter p (p' -ᵥ p) = - affine_map.line_map p p' (-2 * ⟪p' -ᵥ p, p -ᵥ s.center⟫ / ⟪p' -ᵥ p, p' -ᵥ p⟫) := -rfl - -/-- If the vector passed to `second_inter` is given by a subtraction involving the point in -`second_inter`, the result lies in the span of the two points. -/ -lemma sphere.second_inter_vsub_mem_affine_span (s : sphere P) (p₁ p₂ : P) : - s.second_inter p₁ (p₂ -ᵥ p₁) ∈ line[ℝ, p₁, p₂] := -smul_vsub_vadd_mem_affine_span_pair _ _ _ - -/-- If the vector passed to `second_inter` is given by a subtraction involving the point in -`second_inter`, the three points are collinear. -/ -lemma sphere.second_inter_collinear (s : sphere P) (p p' : P) : - collinear ℝ ({p, p', s.second_inter p (p' -ᵥ p)} : set P) := -begin - rw [set.pair_comm, set.insert_comm], - exact (collinear_insert_iff_of_mem_affine_span (s.second_inter_vsub_mem_affine_span _ _)).2 - (collinear_pair ℝ _ _) -end - -/-- If the vector passed to `second_inter` is given by a subtraction involving the point in -`second_inter`, and the second point is not outside the sphere, the second point is weakly -between the first point and the result of `second_inter`. -/ -lemma sphere.wbtw_second_inter {s : sphere P} {p p' : P} (hp : p ∈ s) - (hp' : dist p' s.center ≤ s.radius) : wbtw ℝ p p' (s.second_inter p (p' -ᵥ p)) := -begin - by_cases h : p' = p, { simp [h] }, - refine wbtw_of_collinear_of_dist_center_le_radius (s.second_inter_collinear p p') - hp hp' ((sphere.second_inter_mem _).2 hp) _, - intro he, - rw [eq_comm, sphere.second_inter_eq_self_iff, ←neg_neg (p' -ᵥ p), inner_neg_left, - neg_vsub_eq_vsub_rev, neg_eq_zero, eq_comm] at he, - exact ((inner_pos_or_eq_of_dist_le_radius hp hp').resolve_right (ne.symm h)).ne he -end - -/-- If the vector passed to `second_inter` is given by a subtraction involving the point in -`second_inter`, and the second point is inside the sphere, the second point is strictly between -the first point and the result of `second_inter`. -/ -lemma sphere.sbtw_second_inter {s : sphere P} {p p' : P} (hp : p ∈ s) - (hp' : dist p' s.center < s.radius) : sbtw ℝ p p' (s.second_inter p (p' -ᵥ p)) := -begin - refine ⟨sphere.wbtw_second_inter hp hp'.le, _, _⟩, - { rintro rfl, rw mem_sphere at hp, simpa [hp] using hp' }, - { rintro h, - rw [h, mem_sphere.1 ((sphere.second_inter_mem _).2 hp)] at hp', - exact lt_irrefl _ hp' } -end - -/-- A set of points is concyclic if it is cospherical and coplanar. (Most results are stated -directly in terms of `cospherical` instead of using `concyclic`.) -/ -structure concyclic (ps : set P) : Prop := -(cospherical : cospherical ps) -(coplanar : coplanar ℝ ps) - -/-- A subset of a concyclic set is concyclic. -/ -lemma concyclic.subset {ps₁ ps₂ : set P} (hs : ps₁ ⊆ ps₂) (h : concyclic ps₂) : concyclic ps₁ := -⟨h.1.subset hs, h.2.subset hs⟩ - -/-- The empty set is concyclic. -/ -lemma concyclic_empty : concyclic (∅ : set P) := -⟨cospherical_empty, coplanar_empty ℝ P⟩ - -/-- A single point is concyclic. -/ -lemma concyclic_singleton (p : P) : concyclic ({p} : set P) := -⟨cospherical_singleton p, coplanar_singleton ℝ p⟩ - -/-- Two points are concyclic. -/ -lemma concyclic_pair (p₁ p₂ : P) : concyclic ({p₁, p₂} : set P) := -⟨cospherical_pair p₁ p₂, coplanar_pair ℝ p₁ p₂⟩ - end euclidean_geometry diff --git a/src/geometry/euclidean/circumcenter.lean b/src/geometry/euclidean/circumcenter.lean index da7ffb6e7c325..b5250e2d515f3 100644 --- a/src/geometry/euclidean/circumcenter.lean +++ b/src/geometry/euclidean/circumcenter.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 Joseph Myers. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Joseph Myers -/ -import geometry.euclidean.basic +import geometry.euclidean.sphere.basic import linear_algebra.affine_space.finite_dimensional import tactic.derive_fintype diff --git a/src/geometry/euclidean/sphere/basic.lean b/src/geometry/euclidean/sphere/basic.lean new file mode 100644 index 0000000000000..6eaf3054a2cea --- /dev/null +++ b/src/geometry/euclidean/sphere/basic.lean @@ -0,0 +1,349 @@ +/- +Copyright (c) 2020 Joseph Myers. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joseph Myers +-/ +import analysis.convex.strict_convex_between +import geometry.euclidean.basic + +/-! +# Spheres + +This file defines and proves basic results about spheres and cospherical sets of points in +Euclidean affine spaces. + +## Main definitions + +* `euclidean_geometry.sphere` bundles a `center` and a `radius`. + +* `euclidean_geometry.cospherical` is the property of a set of points being equidistant from some + point. + +* `euclidean_geometry.concyclic` is the property of a set of points being cospherical and + coplanar. + +-/ + +noncomputable theory +open_locale real_inner_product_space + +namespace euclidean_geometry + +variables {V : Type*} (P : Type*) [inner_product_space ℝ V] [metric_space P] [normed_add_torsor V P] + +open finite_dimensional + +/-- A `sphere P` bundles a `center` and `radius`. This definition does not require the radius to +be positive; that should be given as a hypothesis to lemmas that require it. -/ +@[ext] structure sphere := +(center : P) +(radius : ℝ) + +variables {P} + +instance [nonempty P] : nonempty (sphere P) := ⟨⟨classical.arbitrary P, 0⟩⟩ + +instance : has_coe (sphere P) (set P) := ⟨λ s, metric.sphere s.center s.radius⟩ +instance : has_mem P (sphere P) := ⟨λ p s, p ∈ (s : set P)⟩ + +lemma sphere.mk_center (c : P) (r : ℝ) : (⟨c, r⟩ : sphere P).center = c := rfl + +lemma sphere.mk_radius (c : P) (r : ℝ) : (⟨c, r⟩ : sphere P).radius = r := rfl + +@[simp] lemma sphere.mk_center_radius (s : sphere P) : (⟨s.center, s.radius⟩ : sphere P) = s := +by ext; refl + +lemma sphere.coe_def (s : sphere P) : (s : set P) = metric.sphere s.center s.radius := rfl + +@[simp] lemma sphere.coe_mk (c : P) (r : ℝ) : ↑(⟨c, r⟩ : sphere P) = metric.sphere c r := rfl + +@[simp] lemma sphere.mem_coe {p : P} {s : sphere P} : p ∈ (s : set P) ↔ p ∈ s := iff.rfl + +lemma mem_sphere {p : P} {s : sphere P} : p ∈ s ↔ dist p s.center = s.radius := iff.rfl + +lemma mem_sphere' {p : P} {s : sphere P} : p ∈ s ↔ dist s.center p = s.radius := +metric.mem_sphere' + +lemma subset_sphere {ps : set P} {s : sphere P} : ps ⊆ s ↔ ∀ p ∈ ps, p ∈ s := iff.rfl + +lemma dist_of_mem_subset_sphere {p : P} {ps : set P} {s : sphere P} (hp : p ∈ ps) + (hps : ps ⊆ (s : set P)) : dist p s.center = s.radius := +mem_sphere.1 (sphere.mem_coe.1 (set.mem_of_mem_of_subset hp hps)) + +lemma dist_of_mem_subset_mk_sphere {p c : P} {ps : set P} {r : ℝ} (hp : p ∈ ps) + (hps : ps ⊆ ↑(⟨c, r⟩ : sphere P)) : dist p c = r := +dist_of_mem_subset_sphere hp hps + +lemma sphere.ne_iff {s₁ s₂ : sphere P} : + s₁ ≠ s₂ ↔ s₁.center ≠ s₂.center ∨ s₁.radius ≠ s₂.radius := +by rw [←not_and_distrib, ←sphere.ext_iff] + +lemma sphere.center_eq_iff_eq_of_mem {s₁ s₂ : sphere P} {p : P} (hs₁ : p ∈ s₁) (hs₂ : p ∈ s₂) : + s₁.center = s₂.center ↔ s₁ = s₂ := +begin + refine ⟨λ h, sphere.ext _ _ h _, λ h, h ▸ rfl⟩, + rw mem_sphere at hs₁ hs₂, + rw [←hs₁, ←hs₂, h] +end + +lemma sphere.center_ne_iff_ne_of_mem {s₁ s₂ : sphere P} {p : P} (hs₁ : p ∈ s₁) (hs₂ : p ∈ s₂) : + s₁.center ≠ s₂.center ↔ s₁ ≠ s₂ := +(sphere.center_eq_iff_eq_of_mem hs₁ hs₂).not + +lemma dist_center_eq_dist_center_of_mem_sphere {p₁ p₂ : P} {s : sphere P} (hp₁ : p₁ ∈ s) + (hp₂ : p₂ ∈ s) : dist p₁ s.center = dist p₂ s.center := +by rw [mem_sphere.1 hp₁, mem_sphere.1 hp₂] + +lemma dist_center_eq_dist_center_of_mem_sphere' {p₁ p₂ : P} {s : sphere P} (hp₁ : p₁ ∈ s) + (hp₂ : p₂ ∈ s) : dist s.center p₁ = dist s.center p₂ := +by rw [mem_sphere'.1 hp₁, mem_sphere'.1 hp₂] + +/-- A set of points is cospherical if they are equidistant from some +point. In two dimensions, this is the same thing as being +concyclic. -/ +def cospherical (ps : set P) : Prop := +∃ (center : P) (radius : ℝ), ∀ p ∈ ps, dist p center = radius + +/-- The definition of `cospherical`. -/ +lemma cospherical_def (ps : set P) : + cospherical ps ↔ ∃ (center : P) (radius : ℝ), ∀ p ∈ ps, dist p center = radius := +iff.rfl + +/-- A set of points is cospherical if and only if they lie in some sphere. -/ +lemma cospherical_iff_exists_sphere {ps : set P} : + cospherical ps ↔ ∃ s : sphere P, ps ⊆ (s : set P) := +begin + refine ⟨λ h, _, λ h, _⟩, + { rcases h with ⟨c, r, h⟩, + exact ⟨⟨c, r⟩, h⟩ }, + { rcases h with ⟨s, h⟩, + exact ⟨s.center, s.radius, h⟩ } +end + +/-- The set of points in a sphere is cospherical. -/ +lemma sphere.cospherical (s : sphere P) : cospherical (s : set P) := +cospherical_iff_exists_sphere.2 ⟨s, set.subset.rfl⟩ + +/-- A subset of a cospherical set is cospherical. -/ +lemma cospherical.subset {ps₁ ps₂ : set P} (hs : ps₁ ⊆ ps₂) (hc : cospherical ps₂) : + cospherical ps₁ := +begin + rcases hc with ⟨c, r, hcr⟩, + exact ⟨c, r, λ p hp, hcr p (hs hp)⟩ +end + +include V + +/-- The empty set is cospherical. -/ +lemma cospherical_empty : cospherical (∅ : set P) := +begin + use add_torsor.nonempty.some, + simp, +end + +omit V + +/-- A single point is cospherical. -/ +lemma cospherical_singleton (p : P) : cospherical ({p} : set P) := +begin + use p, + simp +end + +include V + +/-- Two points are cospherical. -/ +lemma cospherical_pair (p₁ p₂ : P) : cospherical ({p₁, p₂} : set P) := +begin + use [(2⁻¹ : ℝ) • (p₂ -ᵥ p₁) +ᵥ p₁, (2⁻¹ : ℝ) * (dist p₂ p₁)], + intro p, + rw [set.mem_insert_iff, set.mem_singleton_iff], + rintro ⟨_|_⟩, + { rw [dist_eq_norm_vsub V p₁, vsub_vadd_eq_vsub_sub, vsub_self, zero_sub, norm_neg, norm_smul, + dist_eq_norm_vsub V p₂], + simp }, + { rw [H, dist_eq_norm_vsub V p₂, vsub_vadd_eq_vsub_sub, dist_eq_norm_vsub V p₂], + conv_lhs { congr, congr, rw ←one_smul ℝ (p₂ -ᵥ p₁ : V) }, + rw [←sub_smul, norm_smul], + norm_num } +end + +/-- Any three points in a cospherical set are affinely independent. -/ +lemma cospherical.affine_independent {s : set P} (hs : cospherical s) {p : fin 3 → P} + (hps : set.range p ⊆ s) (hpi : function.injective p) : + affine_independent ℝ p := +begin + rw affine_independent_iff_not_collinear, + intro hc, + rw collinear_iff_of_mem (set.mem_range_self (0 : fin 3)) at hc, + rcases hc with ⟨v, hv⟩, + rw set.forall_range_iff at hv, + have hv0 : v ≠ 0, + { intro h, + have he : p 1 = p 0, by simpa [h] using hv 1, + exact (dec_trivial : (1 : fin 3) ≠ 0) (hpi he) }, + rcases hs with ⟨c, r, hs⟩, + have hs' := λ i, hs (p i) (set.mem_of_mem_of_subset (set.mem_range_self _) hps), + choose f hf using hv, + have hsd : ∀ i, dist ((f i • v) +ᵥ p 0) c = r, + { intro i, + rw ←hf, + exact hs' i }, + have hf0 : f 0 = 0, + { have hf0' := hf 0, + rw [eq_comm, ←@vsub_eq_zero_iff_eq V, vadd_vsub, smul_eq_zero] at hf0', + simpa [hv0] using hf0' }, + have hfi : function.injective f, + { intros i j h, + have hi := hf i, + rw [h, ←hf j] at hi, + exact hpi hi }, + simp_rw [←hsd 0, hf0, zero_smul, zero_vadd, dist_smul_vadd_eq_dist (p 0) c hv0] at hsd, + have hfn0 : ∀ i, i ≠ 0 → f i ≠ 0 := λ i, (hfi.ne_iff' hf0).2, + have hfn0' : ∀ i, i ≠ 0 → f i = (-2) * ⟪v, (p 0 -ᵥ c)⟫ / ⟪v, v⟫, + { intros i hi, + have hsdi := hsd i, + simpa [hfn0, hi] using hsdi }, + have hf12 : f 1 = f 2, { rw [hfn0' 1 dec_trivial, hfn0' 2 dec_trivial] }, + exact (dec_trivial : (1 : fin 3) ≠ 2) (hfi hf12) +end + +/-- Any three points in a cospherical set are affinely independent. -/ +lemma cospherical.affine_independent_of_mem_of_ne {s : set P} (hs : cospherical s) {p₁ p₂ p₃ : P} + (h₁ : p₁ ∈ s) (h₂ : p₂ ∈ s) (h₃ : p₃ ∈ s) (h₁₂ : p₁ ≠ p₂) (h₁₃ : p₁ ≠ p₃) (h₂₃ : p₂ ≠ p₃) : + affine_independent ℝ ![p₁, p₂, p₃] := +begin + refine hs.affine_independent _ _, + { simp [h₁, h₂, h₃, set.insert_subset] }, + { erw [fin.cons_injective_iff, fin.cons_injective_iff], + simp [h₁₂, h₁₃, h₂₃, function.injective] } +end + +/-- The three points of a cospherical set are affinely independent. -/ +lemma cospherical.affine_independent_of_ne {p₁ p₂ p₃ : P} (hs : cospherical ({p₁, p₂, p₃} : set P)) + (h₁₂ : p₁ ≠ p₂) (h₁₃ : p₁ ≠ p₃) (h₂₃ : p₂ ≠ p₃) : + affine_independent ℝ ![p₁, p₂, p₃] := +hs.affine_independent_of_mem_of_ne (set.mem_insert _ _) + (set.mem_insert_of_mem _ (set.mem_insert _ _)) + (set.mem_insert_of_mem _ (set.mem_insert_of_mem _ (set.mem_singleton _))) h₁₂ h₁₃ h₂₃ + +/-- Suppose that `p₁` and `p₂` lie in spheres `s₁` and `s₂`. Then the vector between the centers +of those spheres is orthogonal to that between `p₁` and `p₂`; this is a version of +`inner_vsub_vsub_of_dist_eq_of_dist_eq` for bundled spheres. (In two dimensions, this says that +the diagonals of a kite are orthogonal.) -/ +lemma inner_vsub_vsub_of_mem_sphere_of_mem_sphere {p₁ p₂ : P} {s₁ s₂ : sphere P} + (hp₁s₁ : p₁ ∈ s₁) (hp₂s₁ : p₂ ∈ s₁) (hp₁s₂ : p₁ ∈ s₂) (hp₂s₂ : p₂ ∈ s₂) : + ⟪s₂.center -ᵥ s₁.center, p₂ -ᵥ p₁⟫ = 0 := +inner_vsub_vsub_of_dist_eq_of_dist_eq (dist_center_eq_dist_center_of_mem_sphere hp₁s₁ hp₂s₁) + (dist_center_eq_dist_center_of_mem_sphere hp₁s₂ hp₂s₂) + +/-- Two spheres intersect in at most two points in a two-dimensional subspace containing their +centers; this is a version of `eq_of_dist_eq_of_dist_eq_of_mem_of_finrank_eq_two` for bundled +spheres. -/ +lemma eq_of_mem_sphere_of_mem_sphere_of_mem_of_finrank_eq_two {s : affine_subspace ℝ P} + [finite_dimensional ℝ s.direction] (hd : finrank ℝ s.direction = 2) {s₁ s₂ : sphere P} + {p₁ p₂ p : P} (hs₁ : s₁.center ∈ s) (hs₂ : s₂.center ∈ s) (hp₁s : p₁ ∈ s) (hp₂s : p₂ ∈ s) + (hps : p ∈ s) (hs : s₁ ≠ s₂) (hp : p₁ ≠ p₂) (hp₁s₁ : p₁ ∈ s₁) (hp₂s₁ : p₂ ∈ s₁) (hps₁ : p ∈ s₁) + (hp₁s₂ : p₁ ∈ s₂) (hp₂s₂ : p₂ ∈ s₂) (hps₂ : p ∈ s₂) : p = p₁ ∨ p = p₂ := +eq_of_dist_eq_of_dist_eq_of_mem_of_finrank_eq_two hd hs₁ hs₂ hp₁s hp₂s hps + ((sphere.center_ne_iff_ne_of_mem hps₁ hps₂).2 hs) hp hp₁s₁ hp₂s₁ hps₁ hp₁s₂ hp₂s₂ hps₂ + +/-- Two spheres intersect in at most two points in two-dimensional space; this is a version of +`eq_of_dist_eq_of_dist_eq_of_finrank_eq_two` for bundled spheres. -/ +lemma eq_of_mem_sphere_of_mem_sphere_of_finrank_eq_two [finite_dimensional ℝ V] + (hd : finrank ℝ V = 2) {s₁ s₂ : sphere P} {p₁ p₂ p : P} (hs : s₁ ≠ s₂) (hp : p₁ ≠ p₂) + (hp₁s₁ : p₁ ∈ s₁) (hp₂s₁ : p₂ ∈ s₁) (hps₁ : p ∈ s₁) (hp₁s₂ : p₁ ∈ s₂) (hp₂s₂ : p₂ ∈ s₂) + (hps₂ : p ∈ s₂) : p = p₁ ∨ p = p₂ := +eq_of_dist_eq_of_dist_eq_of_finrank_eq_two hd ((sphere.center_ne_iff_ne_of_mem hps₁ hps₂).2 hs) + hp hp₁s₁ hp₂s₁ hps₁ hp₁s₂ hp₂s₂ hps₂ + +/-- Given a point on a sphere and a point not outside it, the inner product between the +difference of those points and the radius vector is positive unless the points are equal. -/ +lemma inner_pos_or_eq_of_dist_le_radius {s : sphere P} {p₁ p₂ : P} (hp₁ : p₁ ∈ s) + (hp₂ : dist p₂ s.center ≤ s.radius) : 0 < ⟪p₁ -ᵥ p₂, p₁ -ᵥ s.center⟫ ∨ p₁ = p₂ := +begin + by_cases h : p₁ = p₂, { exact or.inr h }, + refine or.inl _, + rw mem_sphere at hp₁, + rw [←vsub_sub_vsub_cancel_right p₁ p₂ s.center, inner_sub_left, + real_inner_self_eq_norm_mul_norm/-, ←dist_eq_norm_vsub, hp₁-/, sub_pos], + refine lt_of_le_of_ne + ((real_inner_le_norm _ _).trans (mul_le_mul_of_nonneg_right _ (norm_nonneg _))) + _, + { rwa [←dist_eq_norm_vsub, ←dist_eq_norm_vsub, hp₁] }, + { rcases hp₂.lt_or_eq with hp₂' | hp₂', + { refine ((real_inner_le_norm _ _).trans_lt (mul_lt_mul_of_pos_right _ _)).ne, + { rwa [←hp₁, @dist_eq_norm_vsub V, @dist_eq_norm_vsub V] at hp₂' }, + { rw [norm_pos_iff, vsub_ne_zero], + rintro rfl, + rw ←hp₁ at hp₂', + refine (dist_nonneg.not_lt : ¬dist p₂ s.center < 0) _, + simpa using hp₂' } }, + { rw [←hp₁, @dist_eq_norm_vsub V, @dist_eq_norm_vsub V] at hp₂', + nth_rewrite 0 ←hp₂', + rw [ne.def, inner_eq_norm_mul_iff_real, hp₂', ←sub_eq_zero, ←smul_sub, + vsub_sub_vsub_cancel_right, ←ne.def, smul_ne_zero_iff, vsub_ne_zero, + and_iff_left (ne.symm h), norm_ne_zero_iff, vsub_ne_zero], + rintro rfl, + refine h (eq.symm _), + simpa using hp₂' } } +end + +/-- Given a point on a sphere and a point not outside it, the inner product between the +difference of those points and the radius vector is nonnegative. -/ +lemma inner_nonneg_of_dist_le_radius {s : sphere P} {p₁ p₂ : P} (hp₁ : p₁ ∈ s) + (hp₂ : dist p₂ s.center ≤ s.radius) : 0 ≤ ⟪p₁ -ᵥ p₂, p₁ -ᵥ s.center⟫ := +begin + rcases inner_pos_or_eq_of_dist_le_radius hp₁ hp₂ with h | rfl, + { exact h.le }, + { simp } +end + +/-- Given a point on a sphere and a point inside it, the inner product between the difference of +those points and the radius vector is positive. -/ +lemma inner_pos_of_dist_lt_radius {s : sphere P} {p₁ p₂ : P} (hp₁ : p₁ ∈ s) + (hp₂ : dist p₂ s.center < s.radius) : 0 < ⟪p₁ -ᵥ p₂, p₁ -ᵥ s.center⟫ := +begin + by_cases h : p₁ = p₂, + { rw [h, mem_sphere] at hp₁, + exact false.elim (hp₂.ne hp₁) }, + exact (inner_pos_or_eq_of_dist_le_radius hp₁ hp₂.le).resolve_right h +end + +/-- Given three collinear points, two on a sphere and one not outside it, the one not outside it +is weakly between the other two points. -/ +lemma wbtw_of_collinear_of_dist_center_le_radius {s : sphere P} {p₁ p₂ p₃ : P} + (h : collinear ℝ ({p₁, p₂, p₃} : set P)) (hp₁ : p₁ ∈ s) (hp₂ : dist p₂ s.center ≤ s.radius) + (hp₃ : p₃ ∈ s) (hp₁p₃ : p₁ ≠ p₃) : wbtw ℝ p₁ p₂ p₃ := +h.wbtw_of_dist_eq_of_dist_le hp₁ hp₂ hp₃ hp₁p₃ + +/-- Given three collinear points, two on a sphere and one inside it, the one inside it is +strictly between the other two points. -/ +lemma sbtw_of_collinear_of_dist_center_lt_radius {s : sphere P} {p₁ p₂ p₃ : P} + (h : collinear ℝ ({p₁, p₂, p₃} : set P)) (hp₁ : p₁ ∈ s) (hp₂ : dist p₂ s.center < s.radius) + (hp₃ : p₃ ∈ s) (hp₁p₃ : p₁ ≠ p₃) : sbtw ℝ p₁ p₂ p₃ := +h.sbtw_of_dist_eq_of_dist_lt hp₁ hp₂ hp₃ hp₁p₃ + +/-- A set of points is concyclic if it is cospherical and coplanar. (Most results are stated +directly in terms of `cospherical` instead of using `concyclic`.) -/ +structure concyclic (ps : set P) : Prop := +(cospherical : cospherical ps) +(coplanar : coplanar ℝ ps) + +/-- A subset of a concyclic set is concyclic. -/ +lemma concyclic.subset {ps₁ ps₂ : set P} (hs : ps₁ ⊆ ps₂) (h : concyclic ps₂) : concyclic ps₁ := +⟨h.1.subset hs, h.2.subset hs⟩ + +/-- The empty set is concyclic. -/ +lemma concyclic_empty : concyclic (∅ : set P) := +⟨cospherical_empty, coplanar_empty ℝ P⟩ + +/-- A single point is concyclic. -/ +lemma concyclic_singleton (p : P) : concyclic ({p} : set P) := +⟨cospherical_singleton p, coplanar_singleton ℝ p⟩ + +/-- Two points are concyclic. -/ +lemma concyclic_pair (p₁ p₂ : P) : concyclic ({p₁, p₂} : set P) := +⟨cospherical_pair p₁ p₂, coplanar_pair ℝ p₁ p₂⟩ + +end euclidean_geometry diff --git a/src/geometry/euclidean/sphere.lean b/src/geometry/euclidean/sphere/power.lean similarity index 64% rename from src/geometry/euclidean/sphere.lean rename to src/geometry/euclidean/sphere/power.lean index 938a557c9c451..4a01046f9d995 100644 --- a/src/geometry/euclidean/sphere.lean +++ b/src/geometry/euclidean/sphere/power.lean @@ -3,40 +3,19 @@ Copyright (c) 2021 Manuel Candales. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Manuel Candales, Benjamin Davidson -/ -import geometry.euclidean.basic -import geometry.euclidean.triangle +import geometry.euclidean.angle.unoriented.affine +import geometry.euclidean.sphere.basic /-! -# Spheres +# Power of a point (intersecting chords and secants) -This file proves basic geometrical results about distances and angles -in spheres in real inner product spaces and Euclidean affine spaces. +This file proves basic geometrical results about power of a point (intersecting chords and +secants) in spheres in real inner product spaces and Euclidean affine spaces. ## Main theorems * `mul_dist_eq_mul_dist_of_cospherical_of_angle_eq_pi`: Intersecting Chords Theorem (Freek No. 55). * `mul_dist_eq_mul_dist_of_cospherical_of_angle_eq_zero`: Intersecting Secants Theorem. -* `mul_dist_add_mul_dist_eq_mul_dist_of_cospherical`: Ptolemy’s Theorem (Freek No. 95). - -TODO: The current statement of Ptolemy’s theorem works around the lack of a "cyclic polygon" concept -in mathlib, which is what the theorem statement would naturally use (or two such concepts, since -both a strict version, where all vertices must be distinct, and a weak version, where consecutive -vertices may be equal, would be useful; Ptolemy's theorem should then use the weak one). - -An API needs to be built around that concept, which would include: -- strict cyclic implies weak cyclic, -- weak cyclic and consecutive points distinct implies strict cyclic, -- weak/strict cyclic implies weak/strict cyclic for any subsequence, -- any three points on a sphere are weakly or strictly cyclic according to whether they are distinct, -- any number of points on a sphere intersected with a two-dimensional affine subspace are cyclic in - some order, -- a list of points is cyclic if and only if its reversal is, -- a list of points is cyclic if and only if any cyclic permutation is, while other permutations - are not when the points are distinct, -- a point P where the diagonals of a cyclic polygon cross exists (and is unique) with weak/strict - betweenness depending on weak/strict cyclicity, -- four points on a sphere with such a point P are cyclic in the appropriate order, -and so on. -/ open real @@ -158,27 +137,4 @@ begin exacts [hab (vsub_left_cancel hab₁).symm, hcd (vsub_left_cancel hcd₁).symm], end -/-- **Ptolemy’s Theorem**. -/ -theorem mul_dist_add_mul_dist_eq_mul_dist_of_cospherical {a b c d p : P} - (h : cospherical ({a, b, c, d} : set P)) - (hapc : ∠ a p c = π) (hbpd : ∠ b p d = π) : - dist a b * dist c d + dist b c * dist d a = dist a c * dist b d := -begin - have h' : cospherical ({a, c, b, d} : set P), { rwa set.insert_comm c b {d} }, - have hmul := mul_dist_eq_mul_dist_of_cospherical_of_angle_eq_pi h' hapc hbpd, - have hbp := left_dist_ne_zero_of_angle_eq_pi hbpd, - have h₁ : dist c d = dist c p / dist b p * dist a b, - { rw [dist_mul_of_eq_angle_of_dist_mul b p a c p d, dist_comm a b], - { rw [angle_eq_angle_of_angle_eq_pi_of_angle_eq_pi hbpd hapc, angle_comm] }, - all_goals { field_simp [mul_comm, hmul] } }, - have h₂ : dist d a = dist a p / dist b p * dist b c, - { rw [dist_mul_of_eq_angle_of_dist_mul c p b d p a, dist_comm c b], - { rwa [angle_comm, angle_eq_angle_of_angle_eq_pi_of_angle_eq_pi], rwa angle_comm }, - all_goals { field_simp [mul_comm, hmul] } }, - have h₃ : dist d p = dist a p * dist c p / dist b p, { field_simp [mul_comm, hmul] }, - have h₄ : ∀ x y : ℝ, x * (y * x) = x * x * y := λ x y, by rw [mul_left_comm, mul_comm], - field_simp [h₁, h₂, dist_eq_add_dist_of_angle_eq_pi hbpd, h₃, hbp, dist_comm a b, - h₄, ← sq, dist_sq_mul_dist_add_dist_sq_mul_dist b, hapc], -end - end euclidean_geometry diff --git a/src/geometry/euclidean/sphere/ptolemy.lean b/src/geometry/euclidean/sphere/ptolemy.lean new file mode 100644 index 0000000000000..cf45fc41b4cdc --- /dev/null +++ b/src/geometry/euclidean/sphere/ptolemy.lean @@ -0,0 +1,72 @@ +/- +Copyright (c) 2021 Manuel Candales. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Manuel Candales, Benjamin Davidson +-/ +import geometry.euclidean.sphere.power +import geometry.euclidean.triangle + +/-! +# Ptolemy's theorem + +This file proves Ptolemy's theorem on the lengths of the diagonals and sides of a cyclic +quadrilateral. + +## Main theorems + +* `mul_dist_add_mul_dist_eq_mul_dist_of_cospherical`: Ptolemy’s Theorem (Freek No. 95). + +TODO: The current statement of Ptolemy’s theorem works around the lack of a "cyclic polygon" concept +in mathlib, which is what the theorem statement would naturally use (or two such concepts, since +both a strict version, where all vertices must be distinct, and a weak version, where consecutive +vertices may be equal, would be useful; Ptolemy's theorem should then use the weak one). + +An API needs to be built around that concept, which would include: +- strict cyclic implies weak cyclic, +- weak cyclic and consecutive points distinct implies strict cyclic, +- weak/strict cyclic implies weak/strict cyclic for any subsequence, +- any three points on a sphere are weakly or strictly cyclic according to whether they are distinct, +- any number of points on a sphere intersected with a two-dimensional affine subspace are cyclic in + some order, +- a list of points is cyclic if and only if its reversal is, +- a list of points is cyclic if and only if any cyclic permutation is, while other permutations + are not when the points are distinct, +- a point P where the diagonals of a cyclic polygon cross exists (and is unique) with weak/strict + betweenness depending on weak/strict cyclicity, +- four points on a sphere with such a point P are cyclic in the appropriate order, +and so on. +-/ + +open real +open_locale euclidean_geometry real_inner_product_space real + +namespace euclidean_geometry + +variables {V : Type*} [inner_product_space ℝ V] +variables {P : Type*} [metric_space P] [normed_add_torsor V P] +include V + +/-- **Ptolemy’s Theorem**. -/ +theorem mul_dist_add_mul_dist_eq_mul_dist_of_cospherical {a b c d p : P} + (h : cospherical ({a, b, c, d} : set P)) + (hapc : ∠ a p c = π) (hbpd : ∠ b p d = π) : + dist a b * dist c d + dist b c * dist d a = dist a c * dist b d := +begin + have h' : cospherical ({a, c, b, d} : set P), { rwa set.insert_comm c b {d} }, + have hmul := mul_dist_eq_mul_dist_of_cospherical_of_angle_eq_pi h' hapc hbpd, + have hbp := left_dist_ne_zero_of_angle_eq_pi hbpd, + have h₁ : dist c d = dist c p / dist b p * dist a b, + { rw [dist_mul_of_eq_angle_of_dist_mul b p a c p d, dist_comm a b], + { rw [angle_eq_angle_of_angle_eq_pi_of_angle_eq_pi hbpd hapc, angle_comm] }, + all_goals { field_simp [mul_comm, hmul] } }, + have h₂ : dist d a = dist a p / dist b p * dist b c, + { rw [dist_mul_of_eq_angle_of_dist_mul c p b d p a, dist_comm c b], + { rwa [angle_comm, angle_eq_angle_of_angle_eq_pi_of_angle_eq_pi], rwa angle_comm }, + all_goals { field_simp [mul_comm, hmul] } }, + have h₃ : dist d p = dist a p * dist c p / dist b p, { field_simp [mul_comm, hmul] }, + have h₄ : ∀ x y : ℝ, x * (y * x) = x * x * y := λ x y, by rw [mul_left_comm, mul_comm], + field_simp [h₁, h₂, dist_eq_add_dist_of_angle_eq_pi hbpd, h₃, hbp, dist_comm a b, + h₄, ← sq, dist_sq_mul_dist_add_dist_sq_mul_dist b, hapc], +end + +end euclidean_geometry diff --git a/src/geometry/euclidean/sphere/second_inter.lean b/src/geometry/euclidean/sphere/second_inter.lean new file mode 100644 index 0000000000000..e00f064381610 --- /dev/null +++ b/src/geometry/euclidean/sphere/second_inter.lean @@ -0,0 +1,175 @@ +/- +Copyright (c) 2022 Joseph Myers. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joseph Myers +-/ +import geometry.euclidean.sphere.basic + +/-! +# Second intersection of a sphere and a line + +This file defines and proves basic results about the second intersection of a sphere with a line +through a point on that sphere. + +## Main definitions + +* `euclidean_geometry.sphere.second_inter` is the second intersection of a sphere with a line + through a point on that sphere. + +-/ + +noncomputable theory +open_locale real_inner_product_space + +namespace euclidean_geometry + +variables {V : Type*} {P : Type*} [inner_product_space ℝ V] [metric_space P] [normed_add_torsor V P] +include V + +/-- The second intersection of a sphere with a line through a point on that sphere; that point +if it is the only point of intersection of the line with the sphere. The intended use of this +definition is when `p ∈ s`; the definition does not use `s.radius`, so in general it returns +the second intersection with the sphere through `p` and with center `s.center`. -/ +def sphere.second_inter (s : sphere P) (p : P) (v : V) : P := +(-2 * ⟪v, p -ᵥ s.center⟫ / ⟪v, v⟫) • v +ᵥ p + +/-- The distance between `second_inter` and the center equals the distance between the original +point and the center. -/ +@[simp] lemma sphere.second_inter_dist (s : sphere P) (p : P) (v : V) : + dist (s.second_inter p v) s.center = dist p s.center := +begin + rw sphere.second_inter, + by_cases hv : v = 0, { simp [hv] }, + rw dist_smul_vadd_eq_dist _ _ hv, + exact or.inr rfl +end + +/-- The point given by `second_inter` lies on the sphere. -/ +@[simp] lemma sphere.second_inter_mem {s : sphere P} {p : P} (v : V) : + s.second_inter p v ∈ s ↔ p ∈ s := +by simp_rw [mem_sphere, sphere.second_inter_dist] + +variables (V) + +/-- If the vector is zero, `second_inter` gives the original point. -/ +@[simp] lemma sphere.second_inter_zero (s : sphere P) (p : P) : + s.second_inter p (0 : V) = p := +by simp [sphere.second_inter] + +variables {V} + +/-- The point given by `second_inter` equals the original point if and only if the line is +orthogonal to the radius vector. -/ +lemma sphere.second_inter_eq_self_iff {s : sphere P} {p : P} {v : V} : + s.second_inter p v = p ↔ ⟪v, p -ᵥ s.center⟫ = 0 := +begin + refine ⟨λ hp, _, λ hp, _⟩, + { by_cases hv : v = 0, { simp [hv] }, + rwa [sphere.second_inter, eq_comm, eq_vadd_iff_vsub_eq, vsub_self, eq_comm, smul_eq_zero, + or_iff_left hv, div_eq_zero_iff, inner_self_eq_zero, or_iff_left hv, mul_eq_zero, + or_iff_right (by norm_num : (-2 : ℝ) ≠ 0)] at hp }, + { rw [sphere.second_inter, hp, mul_zero, zero_div, zero_smul, zero_vadd] } +end + +/-- A point on a line through a point on a sphere equals that point or `second_inter`. -/ +lemma sphere.eq_or_eq_second_inter_of_mem_mk'_span_singleton_iff_mem {s : sphere P} {p : P} + (hp : p ∈ s) {v : V} {p' : P} (hp' : p' ∈ affine_subspace.mk' p (ℝ ∙ v)) : + (p' = p ∨ p' = s.second_inter p v) ↔ p' ∈ s := +begin + refine ⟨λ h, _, λ h, _⟩, + { rcases h with h | h, + { rwa h }, + { rwa [h, sphere.second_inter_mem] } }, + { rw [affine_subspace.mem_mk'_iff_vsub_mem, submodule.mem_span_singleton] at hp', + rcases hp' with ⟨r, hr⟩, + rw [eq_comm, ←eq_vadd_iff_vsub_eq] at hr, + subst hr, + by_cases hv : v = 0, { simp [hv] }, + rw sphere.second_inter, + rw mem_sphere at h hp, + rw [←hp, dist_smul_vadd_eq_dist _ _ hv] at h, + rcases h with h | h; + simp [h] } +end + +/-- `second_inter` is unchanged by multiplying the vector by a nonzero real. -/ +@[simp] lemma sphere.second_inter_smul (s : sphere P) (p : P) (v : V) {r : ℝ} + (hr : r ≠ 0) : s.second_inter p (r • v) = s.second_inter p v := +begin + simp_rw [sphere.second_inter, real_inner_smul_left, inner_smul_right, smul_smul, + div_mul_eq_div_div], + rw [mul_comm, ←mul_div_assoc, ←mul_div_assoc, mul_div_cancel_left _ hr, mul_comm, mul_assoc, + mul_div_cancel_left _ hr, mul_comm] +end + +/-- `second_inter` is unchanged by negating the vector. -/ +@[simp] lemma sphere.second_inter_neg (s : sphere P) (p : P) (v : V) : + s.second_inter p (-v) = s.second_inter p v := +by rw [←neg_one_smul ℝ v, s.second_inter_smul p v (by norm_num : (-1 : ℝ) ≠ 0)] + +/-- Applying `second_inter` twice returns the original point. -/ +@[simp] lemma sphere.second_inter_second_inter (s : sphere P) (p : P) (v : V) : + s.second_inter (s.second_inter p v) v = p := +begin + by_cases hv : v = 0, { simp [hv] }, + have hv' : ⟪v, v⟫ ≠ 0 := inner_self_ne_zero.2 hv, + simp only [sphere.second_inter, vadd_vsub_assoc, vadd_vadd, inner_add_right, inner_smul_right, + div_mul_cancel _ hv'], + rw [←@vsub_eq_zero_iff_eq V, vadd_vsub, ←add_smul, ←add_div], + convert zero_smul ℝ _, + convert zero_div _, + ring +end + +/-- If the vector passed to `second_inter` is given by a subtraction involving the point in +`second_inter`, the result of `second_inter` may be expressed using `line_map`. -/ +lemma sphere.second_inter_eq_line_map (s : sphere P) (p p' : P) : + s.second_inter p (p' -ᵥ p) = + affine_map.line_map p p' (-2 * ⟪p' -ᵥ p, p -ᵥ s.center⟫ / ⟪p' -ᵥ p, p' -ᵥ p⟫) := +rfl + +/-- If the vector passed to `second_inter` is given by a subtraction involving the point in +`second_inter`, the result lies in the span of the two points. -/ +lemma sphere.second_inter_vsub_mem_affine_span (s : sphere P) (p₁ p₂ : P) : + s.second_inter p₁ (p₂ -ᵥ p₁) ∈ line[ℝ, p₁, p₂] := +smul_vsub_vadd_mem_affine_span_pair _ _ _ + +/-- If the vector passed to `second_inter` is given by a subtraction involving the point in +`second_inter`, the three points are collinear. -/ +lemma sphere.second_inter_collinear (s : sphere P) (p p' : P) : + collinear ℝ ({p, p', s.second_inter p (p' -ᵥ p)} : set P) := +begin + rw [set.pair_comm, set.insert_comm], + exact (collinear_insert_iff_of_mem_affine_span (s.second_inter_vsub_mem_affine_span _ _)).2 + (collinear_pair ℝ _ _) +end + +/-- If the vector passed to `second_inter` is given by a subtraction involving the point in +`second_inter`, and the second point is not outside the sphere, the second point is weakly +between the first point and the result of `second_inter`. -/ +lemma sphere.wbtw_second_inter {s : sphere P} {p p' : P} (hp : p ∈ s) + (hp' : dist p' s.center ≤ s.radius) : wbtw ℝ p p' (s.second_inter p (p' -ᵥ p)) := +begin + by_cases h : p' = p, { simp [h] }, + refine wbtw_of_collinear_of_dist_center_le_radius (s.second_inter_collinear p p') + hp hp' ((sphere.second_inter_mem _).2 hp) _, + intro he, + rw [eq_comm, sphere.second_inter_eq_self_iff, ←neg_neg (p' -ᵥ p), inner_neg_left, + neg_vsub_eq_vsub_rev, neg_eq_zero, eq_comm] at he, + exact ((inner_pos_or_eq_of_dist_le_radius hp hp').resolve_right (ne.symm h)).ne he +end + +/-- If the vector passed to `second_inter` is given by a subtraction involving the point in +`second_inter`, and the second point is inside the sphere, the second point is strictly between +the first point and the result of `second_inter`. -/ +lemma sphere.sbtw_second_inter {s : sphere P} {p p' : P} (hp : p ∈ s) + (hp' : dist p' s.center < s.radius) : sbtw ℝ p p' (s.second_inter p (p' -ᵥ p)) := +begin + refine ⟨sphere.wbtw_second_inter hp hp'.le, _, _⟩, + { rintro rfl, rw mem_sphere at hp, simpa [hp] using hp' }, + { rintro h, + rw [h, mem_sphere.1 ((sphere.second_inter_mem _).2 hp)] at hp', + exact lt_irrefl _ hp' } +end + +end euclidean_geometry From acb3d204d4ee883eb686f45d486a2a6811a01329 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 16 Mar 2023 13:33:54 +0000 Subject: [PATCH 0638/1291] chore(algebra/order/field/basic): Rename `pow_minus_two_nonneg` (#18591) Rename `pow_minus_two_nonneg` to `zpow_neg_two_nonneg` and move it to `algebra.order.field.power`. --- src/algebra/order/field/basic.lean | 10 +--------- src/algebra/order/field/power.lean | 1 + src/analysis/special_functions/bernstein.lean | 4 ++-- 3 files changed, 4 insertions(+), 11 deletions(-) diff --git a/src/algebra/order/field/basic.lean b/src/algebra/order/field/basic.lean index d9f98d5104da9..f43a92a264d4e 100644 --- a/src/algebra/order/field/basic.lean +++ b/src/algebra/order/field/basic.lean @@ -470,7 +470,7 @@ lemma strict_mono.div_const {β : Type*} [preorder β] {f : β → α} (hf : str by simpa only [div_eq_mul_inv] using hf.mul_const (inv_pos.2 hc) @[priority 100] -- see Note [lower instance priority] -instance linear_ordered_field.to_densely_ordered : densely_ordered α := +instance linear_ordered_semifield.to_densely_ordered : densely_ordered α := { dense := λ a₁ a₂ h, ⟨(a₁ + a₂) / 2, calc a₁ = (a₁ + a₁) / 2 : (add_self_div_two a₁).symm ... < (a₁ + a₂) / 2 : div_lt_div_of_lt zero_lt_two (add_lt_add_left h _), @@ -810,12 +810,4 @@ lemma abs_inv (a : α) : |a⁻¹| = (|a|)⁻¹ := map_inv₀ (abs_hom : α →* lemma abs_div (a b : α) : |a / b| = |a| / |b| := map_div₀ (abs_hom : α →*₀ α) a b lemma abs_one_div (a : α) : |1 / a| = 1 / |a| := by rw [abs_div, abs_one] -lemma pow_minus_two_nonneg : 0 ≤ a^(-2 : ℤ) := -begin - simp only [inv_nonneg, zpow_neg], - change 0 ≤ a ^ ((2 : ℕ) : ℤ), - rw zpow_coe_nat, - apply sq_nonneg, -end - end diff --git a/src/algebra/order/field/power.lean b/src/algebra/order/field/power.lean index 8c265dd4490bd..d694900f7994e 100644 --- a/src/algebra/order/field/power.lean +++ b/src/algebra/order/field/power.lean @@ -98,6 +98,7 @@ lemma zpow_bit0_nonneg (a : α) (n : ℤ) : 0 ≤ a ^ bit0 n := (mul_self_nonneg _).trans_eq $ (zpow_bit0 _ _).symm lemma zpow_two_nonneg (a : α) : 0 ≤ a ^ (2 : ℤ) := zpow_bit0_nonneg _ _ +lemma zpow_neg_two_nonneg (a : α) : 0 ≤ a ^ (-2 : ℤ) := zpow_bit0_nonneg _ (-1) lemma zpow_bit0_pos (h : a ≠ 0) (n : ℤ) : 0 < a ^ bit0 n := (zpow_bit0_nonneg a n).lt_of_ne (zpow_ne_zero _ h).symm diff --git a/src/analysis/special_functions/bernstein.lean b/src/analysis/special_functions/bernstein.lean index cc66603397cbf..43cc7cc4cfd3f 100644 --- a/src/analysis/special_functions/bernstein.lean +++ b/src/analysis/special_functions/bernstein.lean @@ -228,7 +228,7 @@ begin have npos : 0 < (n:ℝ) := by exact_mod_cast npos', -- Two easy inequalities we'll need later: have w₁ : 0 ≤ 2 * ‖f‖ := mul_nonneg (by norm_num) (norm_nonneg f), - have w₂ : 0 ≤ 2 * ‖f‖ * δ^(-2 : ℤ) := mul_nonneg w₁ pow_minus_two_nonneg, + have w₂ : 0 ≤ 2 * ‖f‖ * δ^(-2 : ℤ) := mul_nonneg w₁ (zpow_neg_two_nonneg _), -- As `[0,1]` is compact, it suffices to check the inequality pointwise. rw (continuous_map.norm_lt_iff _ h), intro x, @@ -293,7 +293,7 @@ begin : mul_le_mul_of_nonneg_left (finset.sum_le_univ_sum_of_nonneg (λ k, mul_nonneg - (mul_nonneg pow_minus_two_nonneg (sq_nonneg _)) + (mul_nonneg (zpow_neg_two_nonneg _) (sq_nonneg _)) bernstein_nonneg)) w₁ ... = (2 * ‖f‖) * δ^(-2 : ℤ) * ∑ k : fin (n+1), (x - k/ₙ)^2 * bernstein n k x : by conv_rhs From b3f4f007a962e3787aa0f3b5c7942a1317f7d88e Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 16 Mar 2023 13:33:55 +0000 Subject: [PATCH 0639/1291] chore(algebra/order/nonneg/*): Separate `floor_ring` from `field` (#18596) Move the `archimedean` and `floor_ring` instances out of `algebra.order.nonneg.field` into a new file `algebra.order.nonneg.floor`. --- src/algebra/order/nonneg/field.lean | 32 ++---------------- src/algebra/order/nonneg/floor.lean | 52 +++++++++++++++++++++++++++++ src/data/rat/nnrat.lean | 1 + src/data/real/nnreal.lean | 5 +-- 4 files changed, 59 insertions(+), 31 deletions(-) create mode 100644 src/algebra/order/nonneg/floor.lean diff --git a/src/algebra/order/nonneg/field.lean b/src/algebra/order/nonneg/field.lean index 916655f8be1f5..4d58ba9ed7e50 100644 --- a/src/algebra/order/nonneg/field.lean +++ b/src/algebra/order/nonneg/field.lean @@ -3,10 +3,10 @@ Copyright (c) 2021 Floris van Doorn. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn -/ -import algebra.order.archimedean -import algebra.order.nonneg.ring -import algebra.order.field.inj_surj +import algebra.order.field.basic import algebra.order.field.canonical.defs +import algebra.order.field.inj_surj +import algebra.order.nonneg.ring /-! # Semifield structure on the type of nonnegative elements @@ -71,30 +71,4 @@ instance linear_ordered_comm_group_with_zero [linear_ordered_field α] : linear_ordered_comm_group_with_zero {x : α // 0 ≤ x} := infer_instance -/-! ### Floor -/ - -instance archimedean [ordered_add_comm_monoid α] [archimedean α] : archimedean {x : α // 0 ≤ x} := -⟨λ x y hy, - let ⟨n, hr⟩ := archimedean.arch (x : α) (hy : (0 : α) < y) in - ⟨n, show (x : α) ≤ (n • y : {x : α // 0 ≤ x}), by simp [*, -nsmul_eq_mul, nsmul_coe]⟩⟩ - -instance floor_semiring [ordered_semiring α] [floor_semiring α] : floor_semiring {r : α // 0 ≤ r} := -{ floor := λ a, ⌊(a : α)⌋₊, - ceil := λ a, ⌈(a : α)⌉₊, - floor_of_neg := λ a ha, floor_semiring.floor_of_neg ha, - gc_floor := λ a n ha, begin - refine (floor_semiring.gc_floor (show 0 ≤ (a : α), from ha)).trans _, - rw [←subtype.coe_le_coe, nonneg.coe_nat_cast] - end, - gc_ceil := λ a n, begin - refine (floor_semiring.gc_ceil (a : α) n).trans _, - rw [←subtype.coe_le_coe, nonneg.coe_nat_cast] - end} - -@[norm_cast] lemma nat_floor_coe [ordered_semiring α] [floor_semiring α] (a : {r : α // 0 ≤ r}) : - ⌊(a : α)⌋₊ = ⌊a⌋₊ := rfl - -@[norm_cast] lemma nat_ceil_coe [ordered_semiring α] [floor_semiring α] (a : {r : α // 0 ≤ r}) : - ⌈(a : α)⌉₊ = ⌈a⌉₊ := rfl - end nonneg diff --git a/src/algebra/order/nonneg/floor.lean b/src/algebra/order/nonneg/floor.lean new file mode 100644 index 0000000000000..ad01bc8b5c5ff --- /dev/null +++ b/src/algebra/order/nonneg/floor.lean @@ -0,0 +1,52 @@ +/- +Copyright (c) 2021 Floris van Doorn. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Floris van Doorn +-/ +import algebra.order.nonneg.ring +import algebra.order.archimedean + +/-! +# Nonnegative elements are archimedean + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + +This file defines instances and prove some properties about the nonnegative elements +`{x : α // 0 ≤ x}` of an arbitrary type `α`. + +This is used to derive algebraic structures on `ℝ≥0` and `ℚ≥0` automatically. + +## Main declarations + +* `{x : α // 0 ≤ x}` is a `floor_semiring` if `α` is. +-/ + +namespace nonneg +variables {α : Type*} + +instance archimedean [ordered_add_comm_monoid α] [archimedean α] : archimedean {x : α // 0 ≤ x} := +⟨λ x y hy, + let ⟨n, hr⟩ := archimedean.arch (x : α) (hy : (0 : α) < y) in + ⟨n, show (x : α) ≤ (n • y : {x : α // 0 ≤ x}), by simp [*, -nsmul_eq_mul, nsmul_coe]⟩⟩ + +instance floor_semiring [ordered_semiring α] [floor_semiring α] : floor_semiring {r : α // 0 ≤ r} := +{ floor := λ a, ⌊(a : α)⌋₊, + ceil := λ a, ⌈(a : α)⌉₊, + floor_of_neg := λ a ha, floor_semiring.floor_of_neg ha, + gc_floor := λ a n ha, begin + refine (floor_semiring.gc_floor (show 0 ≤ (a : α), from ha)).trans _, + rw [←subtype.coe_le_coe, nonneg.coe_nat_cast] + end, + gc_ceil := λ a n, begin + refine (floor_semiring.gc_ceil (a : α) n).trans _, + rw [←subtype.coe_le_coe, nonneg.coe_nat_cast] + end} + +@[norm_cast] lemma nat_floor_coe [ordered_semiring α] [floor_semiring α] (a : {r : α // 0 ≤ r}) : + ⌊(a : α)⌋₊ = ⌊a⌋₊ := rfl + +@[norm_cast] lemma nat_ceil_coe [ordered_semiring α] [floor_semiring α] (a : {r : α // 0 ≤ r}) : + ⌈(a : α)⌉₊ = ⌈a⌉₊ := rfl + +end nonneg diff --git a/src/data/rat/nnrat.lean b/src/data/rat/nnrat.lean index 87cc872ebb68e..416382a3289a4 100644 --- a/src/data/rat/nnrat.lean +++ b/src/data/rat/nnrat.lean @@ -5,6 +5,7 @@ Authors: Yaël Dillies, Bhavik Mehta -/ import algebra.algebra.basic import algebra.order.nonneg.field +import algebra.order.nonneg.floor /-! # Nonnegative rationals diff --git a/src/data/real/nnreal.lean b/src/data/real/nnreal.lean index 09a78d5f3819c..bc8e5bf487e36 100644 --- a/src/data/real/nnreal.lean +++ b/src/data/real/nnreal.lean @@ -3,11 +3,12 @@ Copyright (c) 2018 Johan Commelin. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin -/ -import order.conditionally_complete_lattice.group import algebra.algebra.basic -import algebra.order.nonneg.field import algebra.order.field.canonical.basic +import algebra.order.nonneg.field +import algebra.order.nonneg.floor import data.real.pointwise +import order.conditionally_complete_lattice.group import tactic.positivity /-! From 22f577237f96d87d9084104feba75f3deccd4dd5 Mon Sep 17 00:00:00 2001 From: Anatole Dedecker Date: Thu, 16 Mar 2023 17:03:40 +0000 Subject: [PATCH 0640/1291] feat(analysis/inner_product_space/projection): express `orthogonal_projection` using `linear_proj_of_is_compl` (#18243) We currently don't have any link between `orthogonal_projection` and `linear_proj_of_is_compl`, which blocks us from applying general algebraic facts about projection to `orthogonal_projection`. This PR fixes that. Co-authored-by: Eric Wieser --- src/analysis/inner_product_space/projection.lean | 16 ++++++++++++++++ 1 file changed, 16 insertions(+) diff --git a/src/analysis/inner_product_space/projection.lean b/src/analysis/inner_product_space/projection.lean index ddda82f217ec8..2b1852d5fbbac 100644 --- a/src/analysis/inner_product_space/projection.lean +++ b/src/analysis/inner_product_space/projection.lean @@ -773,6 +773,22 @@ lemma orthogonal_projection_mem_subspace_orthogonal_complement_eq_zero orthogonal_projection K v = 0 := by { ext, convert eq_orthogonal_projection_of_mem_orthogonal _ _; simp [hv] } +lemma orthogonal_projection_eq_linear_proj [complete_space K] (x : E) : + orthogonal_projection K x = + K.linear_proj_of_is_compl _ submodule.is_compl_orthogonal_of_complete_space x := +begin + have : is_compl K Kᗮ := submodule.is_compl_orthogonal_of_complete_space, + nth_rewrite 0 [← submodule.linear_proj_add_linear_proj_of_is_compl_eq_self this x], + rw [map_add, orthogonal_projection_mem_subspace_eq_self, + orthogonal_projection_mem_subspace_orthogonal_complement_eq_zero (submodule.coe_mem _), + add_zero] +end + +lemma orthogonal_projection_coe_linear_map_eq_linear_proj [complete_space K] : + (orthogonal_projection K : E →ₗ[𝕜] K) = + K.linear_proj_of_is_compl _ submodule.is_compl_orthogonal_of_complete_space := +linear_map.ext $ orthogonal_projection_eq_linear_proj + /-- The reflection in `K` of an element of `Kᗮ` is its negation. -/ lemma reflection_mem_subspace_orthogonal_complement_eq_neg [complete_space K] {v : E} (hv : v ∈ Kᗮ) : From 34d3797325d202bd7250431275bb871133cdb611 Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Thu, 16 Mar 2023 17:03:41 +0000 Subject: [PATCH 0641/1291] feat(number_theory/modular_forms): Jacobi theta function (#18564) Define Jacobi's theta function, and prove its functional equation using Poisson summation. --- .../complex/upper_half_plane/basic.lean | 19 ++- src/analysis/special_functions/gaussian.lean | 109 +++++++++++++++--- .../modular_forms/jacobi_theta.lean | 94 +++++++++++++++ 3 files changed, 203 insertions(+), 19 deletions(-) create mode 100644 src/number_theory/modular_forms/jacobi_theta.lean diff --git a/src/analysis/complex/upper_half_plane/basic.lean b/src/analysis/complex/upper_half_plane/basic.lean index 520d31d5e4078..b92328ac57922 100644 --- a/src/analysis/complex/upper_half_plane/basic.lean +++ b/src/analysis/complex/upper_half_plane/basic.lean @@ -309,9 +309,22 @@ variables (x : ℝ) (z : ℍ) end real_add_action -@[simp] lemma modular_S_smul (z : ℍ) : modular_group.S • z = mk (-z : ℂ)⁻¹ z.im_inv_neg_coe_pos := +/- these next few lemmas are *not* flagged `@simp` because of the constructors on the RHS; +instead we use the versions with coercions to `ℂ` as simp lemmas instead. -/ +lemma modular_S_smul (z : ℍ) : modular_group.S • z = mk (-z : ℂ)⁻¹ z.im_inv_neg_coe_pos := by { rw special_linear_group_apply, simp [modular_group.S, neg_div, inv_neg], } +lemma modular_T_zpow_smul (z : ℍ) (n : ℤ) : modular_group.T ^ n • z = (n : ℝ) +ᵥ z := +begin + rw [←subtype.coe_inj, coe_vadd, add_comm, special_linear_group_apply, coe_mk, + modular_group.coe_T_zpow], + simp only [of_apply, cons_val_zero, algebra_map.coe_one, complex.of_real_one, one_mul, + cons_val_one, head_cons, algebra_map.coe_zero, zero_mul, zero_add, div_one], +end + +lemma modular_T_smul (z : ℍ) : modular_group.T • z = (1 : ℝ) +ᵥ z := +by simpa only [algebra_map.coe_one] using modular_T_zpow_smul z 1 + lemma exists_SL2_smul_eq_of_apply_zero_one_eq_zero (g : SL(2, ℝ)) (hc : ↑ₘ[ℝ] g 1 0 = 0) : ∃ (u : {x : ℝ // 0 < x}) (v : ℝ), ((•) g : ℍ → ℍ) = (λ z, v +ᵥ z) ∘ (λ z, u • z) := @@ -335,7 +348,9 @@ begin refine ⟨⟨_, mul_self_pos.mpr hc⟩, c * d, a / c, _⟩, ext1 ⟨z, hz⟩, ext1, suffices : (↑a * z + b) / (↑c * z + d) = a / c - (c * d + ↑c * ↑c * z)⁻¹, - { rw special_linear_group_apply, simpa [-neg_add_rev, inv_neg, ← sub_eq_add_neg], }, + { rw special_linear_group_apply, + simpa only [inv_neg, modular_S_smul, subtype.coe_mk, coe_vadd, complex.of_real_mul, + coe_pos_real_smul, complex.real_smul, function.comp_app, complex.of_real_div] }, replace hc : (c : ℂ) ≠ 0, { norm_cast, assumption, }, replace h_denom : ↑c * z + d ≠ 0, { simpa using h_denom ⟨z, hz⟩, }, have h_aux : (c : ℂ) * d + ↑c * ↑c * z ≠ 0, diff --git a/src/analysis/special_functions/gaussian.lean b/src/analysis/special_functions/gaussian.lean index 2c8c165862218..edebd50124498 100644 --- a/src/analysis/special_functions/gaussian.lean +++ b/src/analysis/special_functions/gaussian.lean @@ -10,6 +10,7 @@ import analysis.convex.complex import analysis.normed.group.basic import analysis.complex.cauchy_integral import measure_theory.group.integration +import analysis.fourier.poisson_summation /-! # Gaussian integral @@ -21,25 +22,31 @@ We prove various versions of the formula for the Gaussian integral: * `integral_gaussian_Ioi` and `integral_gaussian_complex_Ioi`: variants for integrals over `Ioi 0`. * `complex.Gamma_one_half_eq`: the formula `Γ (1 / 2) = √π`. -We also prove, more generally, that the Fourier transform of the Gaussian -is another Gaussian: +We also prove, more generally, that the Fourier transform of the Gaussian is another Gaussian: * `integral_cexp_neg_mul_sq_add_const`: for all complex `b` and `c` with `0 < re b` we have `∫ (x : ℝ), exp (-b * (x + c) ^ 2) = (π / b) ^ (1 / 2)`. * `fourier_transform_gaussian`: for all complex `b` and `t` with `0 < re b`, we have `∫ x:ℝ, exp (I * t * x) * exp (-b * x^2) = (π / b) ^ (1 / 2) * exp (-t ^ 2 / (4 * b))`. * `fourier_transform_gaussian_pi`: a variant with `b` and `t` scaled to give a more symmetric - statement, `∫ x:ℝ, exp (2 * π * I * t * x) * exp (-π * b * x^2) = - (1 / b) ^ (1 / 2) * exp (-π * (1 / b) * t ^ 2)`. + statement, and formulated in terms of the Fourier transform operator `𝓕`. + +As an application, in `real.tsum_exp_neg_mul_int_sq` and `complex.tsum_exp_neg_mul_int_sq`, we use +Poisson summation to prove the identity +`∑' (n : ℤ), exp (-π * a * n ^ 2) = 1 / a ^ (1 / 2) * ∑' (n : ℤ), exp (-π / a * n ^ 2)` +for positive real `a`, or complex `a` with positive real part. (See also +`number_theory.modular_forms.jacobi_theta`.) -/ noncomputable theory open real set measure_theory filter asymptotics -open_locale real topology +open_locale real topology fourier_transform + +open complex (hiding exp continuous_exp abs_of_nonneg sq_abs) -open complex (hiding exp continuous_exp abs_of_nonneg) notation `cexp` := complex.exp +notation `rexp` := real.exp lemma exp_neg_mul_sq_is_o_exp_neg {b : ℝ} (hb : 0 < b) : (λ x:ℝ, exp (-b * x^2)) =o[at_top] (λ x:ℝ, exp (-x)) := @@ -511,26 +518,94 @@ begin simp_rw [this, complex.exp_add, integral_mul_left, integral_cexp_neg_mul_sq_add_const hb] end -lemma _root_.fourier_transform_gaussian_pi (hb : 0 < b.re) (t : ℂ) : - ∫ (x : ℝ), cexp (2 * π * I * t * x) * cexp (-π * b * x ^ 2) = - 1 / b ^ (1 / 2 : ℂ) * cexp (-π * (1 / b) * t ^ 2) := +lemma _root_.fourier_transform_gaussian_pi (hb : 0 < b.re) : + 𝓕 (λ x : ℝ, cexp (-π * b * x ^ 2)) = λ t : ℝ, 1 / b ^ (1 / 2 : ℂ) * cexp (-π / b * t ^ 2) := begin + ext1 t, + simp_rw [fourier_integral_eq_integral_exp_smul, smul_eq_mul], have h1 : 0 < re (π * b) := by { rw of_real_mul_re, exact mul_pos pi_pos hb }, have h2 : b ≠ 0 := by { contrapose! hb, rw [hb, zero_re], }, - convert _root_.fourier_transform_gaussian h1 (2 * π * t) using 1, - { congr' 1, - ext1 x, + convert _root_.fourier_transform_gaussian h1 (-2 * π * t) using 1, + { congr' 1 with x:1, congr' 2, - all_goals { ring } }, + all_goals { push_cast, ring } }, { conv_lhs { rw mul_comm }, congr' 2, { field_simp [of_real_ne_zero.mpr pi_ne_zero], ring, }, - { rw [←div_div, div_self (of_real_ne_zero.mpr pi_ne_zero), cpow_def_of_ne_zero h2, - cpow_def_of_ne_zero (one_div_ne_zero h2), one_div, ←complex.exp_neg, ←neg_mul], - congr' 2, - rw [one_div, complex.log_inv], + { rw [←div_div, div_self (of_real_ne_zero.mpr pi_ne_zero), one_div, one_div b, inv_cpow], rw [ne.def, arg_eq_pi_iff, not_and_distrib, not_lt], exact or.inl hb.le } }, end end gaussian_fourier + +section gaussian_poisson +/-! ## Poisson summation applied to the Gaussian -/ + +variables {E : Type*} [normed_add_comm_group E] + +lemma tendsto_rpow_abs_mul_exp_neg_mul_sq_cocompact {a : ℝ} (ha : 0 < a) (s : ℝ) : + tendsto (λ x : ℝ, |x| ^ s * rexp (-a * x ^ 2)) (cocompact ℝ) (𝓝 0) := +begin + conv in (rexp _) { rw ←sq_abs }, + rw [cocompact_eq, ←comap_abs_at_top, + @tendsto_comap'_iff _ _ _ (λ y, y ^ s * rexp (-a * y ^ 2)) _ _ _ + (mem_at_top_sets.mpr ⟨0, λ b hb, ⟨b, abs_of_nonneg hb⟩⟩)], + exact (rpow_mul_exp_neg_mul_sq_is_o_exp_neg ha s).tendsto_zero_of_tendsto + (tendsto_exp_at_bot.comp $ tendsto_id.neg_const_mul_at_top (neg_lt_zero.mpr one_half_pos)), +end + +lemma is_o_exp_neg_mul_sq_cocompact {a : ℂ} (ha : 0 < a.re) (s : ℝ) : + (λ x : ℝ, complex.exp (-a * x ^ 2)) =o[cocompact ℝ] (λ x : ℝ, |x| ^ s) := +begin + rw ←is_o_norm_left, + simp_rw norm_cexp_neg_mul_sq, + apply is_o_of_tendsto', + { refine eventually.filter_mono cocompact_le_cofinite _, + refine (eventually_cofinite_ne 0).mp (eventually_of_forall (λ x hx h, _)), + exact ((rpow_pos_of_pos (abs_pos.mpr hx) _).ne' h).elim }, + { refine tendsto.congr' (eventually.filter_mono cocompact_le_cofinite _) + (tendsto_zero_iff_norm_tendsto_zero.mp $ + tendsto_rpow_abs_mul_exp_neg_mul_sq_cocompact ha (-s)), + refine (eventually_cofinite_ne 0).mp (eventually_of_forall (λ x hx, _)), + rw [norm_mul, norm_of_nonneg (rpow_nonneg_of_nonneg (abs_nonneg _) _), mul_comm, + rpow_neg (abs_nonneg x), div_eq_mul_inv, norm_of_nonneg (exp_pos _).le] }, +end + +lemma complex.tsum_exp_neg_mul_int_sq {a : ℂ} (ha : 0 < a.re) : + ∑' (n : ℤ), cexp (-π * a * n ^ 2) = 1 / a ^ (1 / 2 : ℂ) * ∑' (n : ℤ), cexp (-π / a * n ^ 2) := +begin + let f := λ x : ℝ, cexp (-π * a * x ^ 2), + have h1 : 0 < (↑π * a).re, + { rw [of_real_mul_re], + exact mul_pos pi_pos ha }, + have h2 : 0 < (↑π / a).re, + { rw [div_eq_mul_inv, of_real_mul_re, inv_re], + refine mul_pos pi_pos (div_pos ha $ norm_sq_pos.mpr _), + contrapose! ha, + rw [ha, zero_re] }, + have f_bd : f =O[cocompact ℝ] (λ x, |x| ^ (-2 : ℝ)), + { convert (is_o_exp_neg_mul_sq_cocompact h1 _).is_O, + ext1 x, + dsimp only [f], + congr' 1, + ring }, + have Ff_bd : 𝓕 f =O[cocompact ℝ] (λ x, |x| ^ (-2 : ℝ)), + { rw fourier_transform_gaussian_pi ha, + convert (is_o_exp_neg_mul_sq_cocompact h2 _).is_O.const_mul_left _, + ext1 x, + congr' 1, + ring_nf }, + simpa only [fourier_transform_gaussian_pi ha, tsum_mul_left] using + real.tsum_eq_tsum_fourier_integral_of_rpow_decay + (complex.continuous_exp.comp (continuous_const.mul (continuous_of_real.pow 2)) : continuous f) + one_lt_two f_bd Ff_bd +end + +lemma real.tsum_exp_neg_mul_int_sq {a : ℝ} (ha : 0 < a) : + ∑' (n : ℤ), exp (-π * a * n ^ 2) = 1 / a ^ (1 / 2 : ℝ) * ∑' (n : ℤ), exp (-π / a * n ^ 2) := +by simpa only [←of_real_inj, of_real_mul, of_real_tsum, of_real_exp, of_real_div, of_real_pow, + of_real_int_cast, of_real_neg, of_real_cpow ha.le, of_real_bit0, of_real_one] using + complex.tsum_exp_neg_mul_int_sq (by rwa [of_real_re] : 0 < (a : ℂ).re) + +end gaussian_poisson diff --git a/src/number_theory/modular_forms/jacobi_theta.lean b/src/number_theory/modular_forms/jacobi_theta.lean new file mode 100644 index 0000000000000..99e025059d26f --- /dev/null +++ b/src/number_theory/modular_forms/jacobi_theta.lean @@ -0,0 +1,94 @@ +/- +Copyright (c) 2023 David Loeffler. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: David Loeffler +-/ +import analysis.complex.upper_half_plane.basic +import analysis.special_functions.gaussian + +/-! # Jacobi's theta function + +This file defines the Jacobi theta function + +$$\theta(\tau) = \sum_{n \in \mathbb{Z}} \exp (i \pi n ^ 2 \tau),$$ + +and proves the modular transformation properties `θ (τ + 2) = θ τ` and +`θ (-1 / τ) = (-I * τ) ^ (1 / 2) * θ τ`, using Poisson's summation formula for the latter. +-/ + +open complex real + +open_locale real big_operators upper_half_plane + +/-- Jacobi's theta function `∑' (n : ℤ), exp (π * I * n ^ 2 * τ)`. -/ +noncomputable def jacobi_theta (τ : ℍ) : ℂ := ∑' (n : ℤ), cexp (π * I * n ^ 2 * τ) + +lemma jacobi_theta_unif_summable {R : ℝ} (hR : 0 < R) : + ∃ (bd : ℤ → ℝ), (summable bd) ∧ + (∀ {τ : ℍ} (hτ : R ≤ τ.im) (n : ℤ), ‖cexp (π * I * n ^ 2 * τ)‖ ≤ bd n) := +begin + let y := rexp (-π * R), + have h : y < 1, from exp_lt_one_iff.mpr (mul_neg_of_neg_of_pos (neg_lt_zero.mpr pi_pos) hR), + refine ⟨λ n, y ^ |n|, summable_int_of_summable_nat _ _, λ τ hτ n, _⟩, swap 3, + { refine le_trans _ (_ : y ^ (n ^ 2) ≤ y ^ |n|), + { rw [complex.norm_eq_abs, complex.abs_exp], + have : (↑π * I * n ^ 2 * τ).re = (-π * (τ : ℂ).im) * n ^ 2, + { rw [(by { push_cast, ring } : ↑π * I * n ^ 2 * τ = ↑(π * n ^ 2) * (τ * I)), + of_real_mul_re, mul_I_re], + ring }, + obtain ⟨m, hm⟩ := int.eq_coe_of_zero_le (sq_nonneg n), + rw [this, exp_mul, ←int.cast_pow, rpow_int_cast, hm, zpow_coe_nat, zpow_coe_nat], + refine pow_le_pow_of_le_left (exp_pos _).le (real.exp_le_exp.mpr _) _, + rwa [mul_le_mul_left_of_neg (neg_lt_zero.mpr pi_pos), upper_half_plane.coe_im] }, + { rw [←inv_inv y, inv_zpow' _ (|n|), inv_zpow' _ (n ^ 2)], + refine zpow_le_of_le (one_le_inv (exp_pos _) h.le) (neg_le_neg _), + rw [int.abs_eq_nat_abs, ←int.nat_abs_sq, ←nat.cast_pow, nat.cast_le, sq], + exact n.nat_abs.le_mul_self } }, + all_goals { simp only [abs_neg, int.abs_coe_nat, zpow_coe_nat], + exact summable_geometric_of_lt_1 (real.exp_pos _).le h }, +end + +lemma jacobi_theta_summable (τ : ℍ) : summable (λ n : ℤ, cexp (π * I * n ^ 2 * τ)) := +let ⟨bd, h, h'⟩ := jacobi_theta_unif_summable τ.im_pos in + summable_norm_iff.mp (summable_of_nonneg_of_le (λ n, norm_nonneg _) (h' $ le_refl _) h) + +lemma jacobi_theta_two_vadd (τ : ℍ) : jacobi_theta ((2 : ℝ) +ᵥ τ) = jacobi_theta τ := +begin + refine tsum_congr (λ n, _), + rw [upper_half_plane.coe_vadd, of_real_bit0, of_real_one], + suffices : cexp (↑π * I * ↑n ^ 2 * 2) = 1, by rw [mul_add, complex.exp_add, this, one_mul], + rw [(by { push_cast, ring } : ↑π * I * ↑n ^ 2 * 2 = ↑(n ^ 2) * (2 * π * I)), + complex.exp_int_mul, complex.exp_two_pi_mul_I, one_zpow], +end + +lemma jacobi_theta_T_sq_smul (τ : ℍ) : jacobi_theta (modular_group.T ^ 2 • τ) = jacobi_theta τ := +begin + suffices : (2 : ℝ) +ᵥ τ = modular_group.T ^ (2 : ℤ) • τ, from this ▸ (jacobi_theta_two_vadd τ), + simp only [←subtype.coe_inj, upper_half_plane.modular_T_zpow_smul, int.cast_two], +end + +lemma jacobi_theta_S_smul (τ : ℍ) : + jacobi_theta (modular_group.S • τ) = (-I * τ) ^ (1 / 2 : ℂ) * jacobi_theta τ := +begin + unfold jacobi_theta, + rw [upper_half_plane.modular_S_smul, upper_half_plane.coe_mk], + have ha : 0 < (-I * τ).re, + { rw [neg_mul, neg_re, mul_re, I_re, I_im, zero_mul, one_mul, zero_sub, neg_neg], + exact τ.im_pos }, + have ha' : (-I * τ) ^ (1 / 2 : ℂ) ≠ 0, + { rw [ne.def, cpow_eq_zero_iff], + contrapose! ha, + rw [ha.1, zero_re] }, + have hτ : (τ : ℂ) ≠ 0, from τ.ne_zero, + have := complex.tsum_exp_neg_mul_int_sq ha, + rw [mul_comm ((1:ℂ) / _) _, mul_one_div, eq_div_iff ha', mul_comm _ (_ ^ _), eq_comm] at this, + convert this using 3, + { ext1 n, + congr' 1, + field_simp [hτ, I_ne_zero], + ring_nf, + rw [I_sq, mul_neg, mul_one, neg_mul, neg_neg] }, + { ext1 n, + congr' 1, + ring_nf } +end From acee671f47b8e7972a1eb6f4eed74b4b3abce829 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 16 Mar 2023 17:03:43 +0000 Subject: [PATCH 0642/1291] chore(data/{nat,int}/cast/field): generalize to division_ring (#18598) Notably, this now works on quaternions. Forward-ported at https://github.com/leanprover-community/mathlib4/pull/2928 --- src/algebra/char_zero/lemmas.lean | 2 +- src/data/int/cast/field.lean | 8 ++++---- src/data/int/char_zero.lean | 2 +- src/data/nat/cast/field.lean | 7 ++++--- 4 files changed, 10 insertions(+), 9 deletions(-) diff --git a/src/algebra/char_zero/lemmas.lean b/src/algebra/char_zero/lemmas.lean index ea21ced81e110..37329fd11dfbd 100644 --- a/src/algebra/char_zero/lemmas.lean +++ b/src/algebra/char_zero/lemmas.lean @@ -35,7 +35,7 @@ def cast_embedding : ℕ ↪ R := ⟨coe, cast_injective⟩ by { rw [←cast_pow, cast_eq_one], exact pow_eq_one_iff hn } @[simp, norm_cast] -theorem cast_div_char_zero {k : Type*} [field k] [char_zero k] {m n : ℕ} +theorem cast_div_char_zero {k : Type*} [division_semiring 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, diff --git a/src/data/int/cast/field.lean b/src/data/int/cast/field.lean index 2fe7363a6d4f2..30f364581d4bd 100644 --- a/src/data/int/cast/field.lean +++ b/src/data/int/cast/field.lean @@ -28,18 +28,18 @@ variables {α : Type*} /-- Auxiliary lemma for norm_cast to move the cast `-↑n` upwards to `↑-↑n`. -(The restriction to `field` is necessary, otherwise this would also apply in the case where +(The restriction to `division_ring` is necessary, otherwise this would also apply in the case where `R = ℤ` and cause nontermination.) -/ @[norm_cast] -lemma cast_neg_nat_cast {R} [field R] (n : ℕ) : ((-n : ℤ) : R) = -n := by simp +lemma cast_neg_nat_cast {R} [division_ring R] (n : ℕ) : ((-n : ℤ) : R) = -n := by simp -@[simp] theorem cast_div [field α] {m n : ℤ} (n_dvd : n ∣ m) (n_nonzero : (n : α) ≠ 0) : +@[simp] theorem cast_div [division_ring α] {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], + rw [int.mul_div_cancel_left _ this, mul_comm n k, int.cast_mul, mul_div_cancel _ n_nonzero], end end int diff --git a/src/data/int/char_zero.lean b/src/data/int/char_zero.lean index 8a6cb00890c84..91d70f00d7956 100644 --- a/src/data/int/char_zero.lean +++ b/src/data/int/char_zero.lean @@ -38,7 +38,7 @@ theorem cast_ne_zero [add_group_with_one α] [char_zero α] {n : ℤ} : (n : α) not_congr cast_eq_zero @[simp, norm_cast] -theorem cast_div_char_zero {k : Type*} [field k] [char_zero k] {m n : ℤ} +theorem cast_div_char_zero {k : Type*} [division_ring 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, diff --git a/src/data/nat/cast/field.lean b/src/data/nat/cast/field.lean index 8b316e9f7e19c..5bae4899761ee 100644 --- a/src/data/nat/cast/field.lean +++ b/src/data/nat/cast/field.lean @@ -25,15 +25,16 @@ namespace nat variables {α : Type*} -@[simp] theorem cast_div [field α] {m n : ℕ} (n_dvd : n ∣ m) (n_nonzero : (n : α) ≠ 0) : +@[simp] theorem cast_div [division_semiring α] {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 _ this.bot_lt, cast_mul, mul_div_cancel_left _ n_nonzero], + rw [nat.mul_div_cancel_left _ this.bot_lt, mul_comm n k, cast_mul, mul_div_cancel _ n_nonzero], end -lemma cast_div_div_div_cancel_right [field α] [char_zero α] {m n d : ℕ} (hn : d ∣ n) (hm : d ∣ m) : +lemma cast_div_div_div_cancel_right [division_semiring α] [char_zero α] {m n d : ℕ} + (hn : d ∣ n) (hm : d ∣ m) : (↑(m / d) : α) / (↑(n / d) : α) = (m : α) / n := begin rcases eq_or_ne d 0 with rfl | hd, { simp [zero_dvd_iff.mp hm], }, From da3fc4a33ff6bc75f077f691dc94c217b8d41559 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 16 Mar 2023 20:41:23 +0000 Subject: [PATCH 0643/1291] feat(analysis/normed_space/quaternion_exponential): lemmas about the quaternion exponential (#18349) This gives the result that $\exp q = \exp r (\cos \|v\| + \frac{v}{\|v\|} \sin \|v\|)$ where $r$ is the real part and $v$ the imaginary part. After adding some missing algebraic lemmas, the result that $\|\exp q\| = \|\exp r\|$ is then simple. --- src/algebra/quaternion.lean | 50 ++++++- src/analysis/normed_space/exponential.lean | 17 ++- .../normed_space/quaternion_exponential.lean | 134 ++++++++++++++++++ src/analysis/quaternion.lean | 25 ++++ 4 files changed, 211 insertions(+), 15 deletions(-) create mode 100644 src/analysis/normed_space/quaternion_exponential.lean diff --git a/src/algebra/quaternion.lean b/src/algebra/quaternion.lean index 9939c2410e6b1..c5606fcbeb912 100644 --- a/src/algebra/quaternion.lean +++ b/src/algebra/quaternion.lean @@ -415,13 +415,20 @@ lemma eq_re_iff_mem_range_coe {a : ℍ[R, c₁, c₂]} : a = a.re ↔ a ∈ set.range (coe : R → ℍ[R, c₁, c₂]) := ⟨λ h, ⟨a.re, h.symm⟩, λ ⟨x, h⟩, eq_re_of_eq_coe h.symm⟩ +section char_zero +variables [no_zero_divisors R] [char_zero R] + @[simp] -lemma conj_fixed {R : Type*} [comm_ring R] [no_zero_divisors R] [char_zero R] - {c₁ c₂ : R} {a : ℍ[R, c₁, c₂]} : +lemma conj_eq_self {c₁ c₂ : R} {a : ℍ[R, c₁, c₂]} : conj a = a ↔ a = a.re := by simp [ext_iff, neg_eq_iff_add_eq_zero, add_self_eq_zero] --- Can't use `rw ← conj_fixed` in the proof without additional assumptions +lemma conj_eq_neg {c₁ c₂ : R} {a : ℍ[R, c₁, c₂]} : + conj a = -a ↔ a.re = 0 := +by simp [ext_iff, eq_neg_iff_add_eq_zero] + +end char_zero +-- Can't use `rw ← conj_eq_self` in the proof without additional assumptions lemma conj_mul_eq_coe : conj a * a = (conj a * a).re := by ext; simp; ring_exp @@ -687,9 +694,14 @@ quaternion_algebra.eq_re_of_eq_coe h lemma eq_re_iff_mem_range_coe {a : ℍ[R]} : a = a.re ↔ a ∈ set.range (coe : R → ℍ[R]) := quaternion_algebra.eq_re_iff_mem_range_coe -@[simp] lemma conj_fixed {R : Type*} [comm_ring R] [no_zero_divisors R] [char_zero R] {a : ℍ[R]} : - conj a = a ↔ a = a.re := -quaternion_algebra.conj_fixed +section char_zero +variables [no_zero_divisors R] [char_zero R] + +@[simp] lemma conj_eq_self {a : ℍ[R]} : conj a = a ↔ a = a.re := quaternion_algebra.conj_eq_self + +@[simp] lemma conj_eq_neg {a : ℍ[R]} : conj a = -a ↔ a.re = 0 := quaternion_algebra.conj_eq_neg + +end char_zero lemma conj_mul_eq_coe : conj a * a = (conj a * a).re := a.conj_mul_eq_coe @@ -749,6 +761,16 @@ lemma coe_norm_sq_add : (norm_sq (a + b) : ℍ[R]) = norm_sq a + a * b.conj + b * a.conj + norm_sq b := by simp [← self_mul_conj, mul_add, add_mul, add_assoc] +lemma norm_sq_smul (r : R) (q : ℍ[R]) : norm_sq (r • q) = r^2 * norm_sq q := +by simp_rw [norm_sq_def, conj_smul, smul_mul_smul, smul_re, sq, smul_eq_mul] + +lemma norm_sq_add (a b : ℍ[R]) : norm_sq (a + b) = norm_sq a + norm_sq b + 2 * (a * conj b).re := +calc norm_sq (a + b) = (norm_sq a + (a * conj b).re) + ((b * conj a).re + norm_sq b) + : by simp_rw [norm_sq_def, conj_add, add_mul, mul_add, add_re] + ... = norm_sq a + norm_sq b + ((a * conj b).re + (b * conj a).re) : by abel + ... = norm_sq a + norm_sq b + 2 * (a * conj b).re + : by rw [←add_re, ←conj_mul_conj a b, self_add_conj', coe_re] + end quaternion namespace quaternion @@ -787,6 +809,22 @@ instance : no_zero_divisors ℍ[R] := instance : is_domain ℍ[R] := no_zero_divisors.to_is_domain _ +lemma sq_eq_norm_sq : a^2 = norm_sq a ↔ a = a.re := +begin + simp_rw [←conj_eq_self], + obtain rfl | hq0 := eq_or_ne a 0, + { simp }, + { rw [←conj_mul_self, sq, mul_left_inj' hq0, eq_comm] } +end + +lemma sq_eq_neg_norm_sq : a^2 = -norm_sq a ↔ a.re = 0 := +begin + simp_rw [←conj_eq_neg], + obtain rfl | hq0 := eq_or_ne a 0, + { simp }, + rw [←conj_mul_self, ←mul_neg, ←neg_sq, sq, mul_left_inj' (neg_ne_zero.mpr hq0), eq_comm], +end + end linear_ordered_comm_ring section field diff --git a/src/analysis/normed_space/exponential.lean b/src/analysis/normed_space/exponential.lean index ed56dd6b68828..879be3bb81e5b 100644 --- a/src/analysis/normed_space/exponential.lean +++ b/src/analysis/normed_space/exponential.lean @@ -99,18 +99,17 @@ tsum_congr (λ n, exp_series_apply_eq x n) lemma exp_eq_tsum : exp 𝕂 = (λ x : 𝔸, ∑' (n : ℕ), (n!⁻¹ : 𝕂) • x^n) := funext exp_series_sum_eq -@[simp] lemma exp_zero [t2_space 𝔸] : exp 𝕂 (0 : 𝔸) = 1 := +lemma exp_series_apply_zero (n : ℕ) : exp_series 𝕂 𝔸 n (λ _, (0 : 𝔸)) = pi.single 0 1 n := begin - 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], - simp }, - refine tsum_congr (λ n, _), - split_ifs with h h; - simp [h] + rw exp_series_apply_eq, + cases n, + { rw [pow_zero, nat.factorial_zero, nat.cast_one, inv_one, one_smul, pi.single_eq_same], }, + { rw [zero_pow (nat.succ_pos _), smul_zero, pi.single_eq_of_ne (n.succ_ne_zero)], }, end +@[simp] lemma exp_zero [t2_space 𝔸] : exp 𝕂 (0 : 𝔸) = 1 := +by simp_rw [exp_eq_tsum, ←exp_series_apply_eq, exp_series_apply_zero, tsum_pi_single] + @[simp] lemma exp_op [t2_space 𝔸] (x : 𝔸) : exp 𝕂 (mul_opposite.op x) = mul_opposite.op (exp 𝕂 x) := by simp_rw [exp, exp_series_sum_eq, ←mul_opposite.op_pow, ←mul_opposite.op_smul, tsum_op] diff --git a/src/analysis/normed_space/quaternion_exponential.lean b/src/analysis/normed_space/quaternion_exponential.lean new file mode 100644 index 0000000000000..98345379d5517 --- /dev/null +++ b/src/analysis/normed_space/quaternion_exponential.lean @@ -0,0 +1,134 @@ +/- +Copyright (c) 2023 Eric Wieser. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Eric Wieser +-/ +import analysis.quaternion +import analysis.normed_space.exponential +import analysis.inner_product_space.pi_L2 +import analysis.special_functions.trigonometric.series + +/-! +# Lemmas about `exp` on `quaternion`s + +This file contains results about `exp` on `quaternion ℝ`. + +## Main results + +* `quaternion.exp_eq`: the general expansion of the quaternion exponential in terms of `real.cos` + and `real.sin`. +* `quaternion.exp_of_re_eq_zero`: the special case when the quaternion has a zero real part. +* `quaternion.norm_exp`: the norm of the quaternion exponential is the norm of the exponential of + the real part. + +-/ + +open_locale quaternion nat + +namespace quaternion + +lemma conj_exp (q : ℍ[ℝ]) : conj (exp ℝ q) = exp ℝ (conj q) := star_exp q + +@[simp, norm_cast] lemma exp_coe (r : ℝ) : exp ℝ (r : ℍ[ℝ]) = ↑(exp ℝ r) := +(map_exp ℝ (algebra_map ℝ ℍ[ℝ]) (continuous_algebra_map _ _) _).symm + +/-- Auxiliary result; if the power series corresponding to `real.cos` and `real.sin` evaluated +at `‖q‖` tend to `c` and `s`, then the exponential series tends to `c + (s / ‖q‖)`. -/ +lemma has_sum_exp_series_of_imaginary + {q : quaternion ℝ} (hq : q.re = 0) {c s : ℝ} + (hc : has_sum (λ n, (-1)^n * ‖q‖^(2 * n) / (2 * n)!) c) + (hs : has_sum (λ n, (-1)^n * ‖q‖^(2 * n + 1) / (2 * n + 1)!) s) : + has_sum (λ n, exp_series ℝ _ n (λ _, q)) (↑c + (s / ‖q‖) • q) := +begin + replace hc := has_sum_coe.mpr hc, + replace hs := (hs.div_const ‖q‖).smul_const q, + obtain rfl | hq0 := eq_or_ne q 0, + { simp_rw [exp_series_apply_zero, norm_zero, div_zero, zero_smul, add_zero], + simp_rw [norm_zero] at hc, + convert hc, + ext (_ | n) : 1, + { rw [pow_zero, mul_zero, pow_zero, nat.factorial_zero, nat.cast_one, div_one, one_mul, + pi.single_eq_same, coe_one], }, + { rw [zero_pow (mul_pos two_pos (nat.succ_pos _)), mul_zero, zero_div, + pi.single_eq_of_ne (n.succ_ne_zero), coe_zero], } }, + simp_rw exp_series_apply_eq, + have hq2 : q^2 = -norm_sq q := sq_eq_neg_norm_sq.mpr hq, + have hqn := norm_ne_zero_iff.mpr hq0, + refine has_sum.even_add_odd _ _, + { convert hc using 1, + ext n : 1, + let k : ℝ := ↑(2 * n)!, + calc k⁻¹ • q ^ (2 * n) + = k⁻¹ • ((-norm_sq q) ^ n) : by rw [pow_mul, hq2] + ... = k⁻¹ • ↑((-1) ^ n * ‖q‖ ^ (2 * n)) : _ + ... = ↑((-1) ^ n * ‖q‖ ^ (2 * n) / k) : _, + { congr' 1, + rw [neg_pow, norm_sq_eq_norm_sq, pow_mul, sq], + push_cast }, + { rw [←coe_mul_eq_smul, div_eq_mul_inv], + norm_cast, + ring_nf } }, + { convert hs using 1, + ext n : 1, + let k : ℝ := ↑(2 * n + 1)!, + calc k⁻¹ • q ^ (2 * n + 1) + = k⁻¹ • ((-norm_sq q) ^ n * q) : by rw [pow_succ', pow_mul, hq2] + ... = k⁻¹ • ((-1) ^ n * ‖q‖ ^ (2 * n)) • q : _ + ... = ((-1) ^ n * ‖q‖ ^ (2 * n + 1) / k / ‖q‖) • q : _, + { congr' 1, + rw [neg_pow, norm_sq_eq_norm_sq, pow_mul, sq, ←coe_mul_eq_smul], + push_cast }, + { rw smul_smul, + congr' 1, + simp_rw [pow_succ', mul_div_assoc, div_div_cancel_left' hqn], + ring } }, +end + +/-- The closed form for the quaternion exponential on imaginary quaternions. -/ +lemma exp_of_re_eq_zero (q : quaternion ℝ) (hq : q.re = 0) : + exp ℝ q = ↑(real.cos ‖q‖) + (real.sin ‖q‖ / ‖q‖) • q := +begin + rw [exp_eq_tsum], + refine has_sum.tsum_eq _, + simp_rw ← exp_series_apply_eq, + exact has_sum_exp_series_of_imaginary hq (real.has_sum_cos _) (real.has_sum_sin _), +end + +/-- The closed form for the quaternion exponential on arbitrary quaternions. -/ +lemma exp_eq (q : quaternion ℝ) : + exp ℝ q = exp ℝ q.re • (↑(real.cos ‖q.im‖) + (real.sin ‖q.im‖ / ‖q.im‖) • q.im) := +begin + rw [←exp_of_re_eq_zero q.im q.im_re, ←coe_mul_eq_smul, ←exp_coe, ←exp_add_of_commute, re_add_im], + exact algebra.commutes q.re (_ : ℍ[ℝ]), +end + +lemma re_exp (q : ℍ[ℝ]) : (exp ℝ q).re = exp ℝ q.re * (real.cos ‖q - q.re‖) := +by simp [exp_eq] + +lemma im_exp (q : ℍ[ℝ]) : (exp ℝ q).im = (exp ℝ q.re * (real.sin ‖q.im‖ / ‖q.im‖)) • q.im := +by simp [exp_eq, smul_smul] + +lemma norm_sq_exp (q : ℍ[ℝ]) : norm_sq (exp ℝ q) = (exp ℝ q.re)^2 := +calc norm_sq (exp ℝ q) + = norm_sq (exp ℝ q.re • (↑(real.cos ‖q.im‖) + (real.sin ‖q.im‖ / ‖q.im‖) • q.im)) + : by rw exp_eq +... = (exp ℝ q.re)^2 * norm_sq ((↑(real.cos ‖q.im‖) + (real.sin ‖q.im‖ / ‖q.im‖) • q.im)) + : by rw [norm_sq_smul] +... = (exp ℝ q.re)^2 * ((real.cos ‖q.im‖) ^ 2 + (real.sin ‖q.im‖)^2) + : begin + congr' 1, + obtain hv | hv := eq_or_ne (‖q.im‖) 0, + { simp [hv] }, + rw [norm_sq_add, norm_sq_smul, conj_smul, coe_mul_eq_smul, smul_re, smul_re, conj_re, im_re, + smul_zero, smul_zero, mul_zero, add_zero, div_pow, norm_sq_coe, norm_sq_eq_norm_sq, ←sq, + div_mul_cancel _ (pow_ne_zero _ hv)], + end +... = (exp ℝ q.re)^2 : by rw [real.cos_sq_add_sin_sq, mul_one] + +/-- Note that this implies that exponentials of pure imaginary quaternions are unit quaternions +since in that case the RHS is `1` via `exp_zero` and `norm_one`. -/ +@[simp] lemma norm_exp (q : ℍ[ℝ]) : ‖exp ℝ q‖ = ‖exp ℝ q.re‖ := +by rw [norm_eq_sqrt_real_inner (exp ℝ q), inner_self, norm_sq_exp, real.sqrt_sq_eq_abs, + real.norm_eq_abs] + +end quaternion diff --git a/src/analysis/quaternion.lean b/src/analysis/quaternion.lean index 85c7306260f1a..762cb2d8533d1 100644 --- a/src/analysis/quaternion.lean +++ b/src/analysis/quaternion.lean @@ -157,4 +157,29 @@ begin exact (complete_space_congr this).1 (by apply_instance) end +section infinite_sum +variables {α : Type*} + +@[simp, norm_cast] lemma has_sum_coe {f : α → ℝ} {r : ℝ} : + has_sum (λ a, (f a : ℍ)) (↑r : ℍ) ↔ has_sum f r := +⟨λ h, by simpa only using + h.map (show ℍ →ₗ[ℝ] ℝ, from quaternion_algebra.re_lm _ _) continuous_re, + λ h, by simpa only using h.map (algebra_map ℝ ℍ) (continuous_algebra_map _ _)⟩ + +@[simp, norm_cast] +lemma summable_coe {f : α → ℝ} : summable (λ a, (f a : ℍ)) ↔ summable f := +by simpa only using summable.map_iff_of_left_inverse (algebra_map ℝ ℍ) + (show ℍ →ₗ[ℝ] ℝ, from quaternion_algebra.re_lm _ _) + (continuous_algebra_map _ _) continuous_re coe_re + +@[norm_cast] lemma tsum_coe (f : α → ℝ) : ∑' a, (f a : ℍ) = ↑(∑' a, f a) := +begin + by_cases hf : summable f, + { exact (has_sum_coe.mpr hf.has_sum).tsum_eq, }, + { simp [tsum_eq_zero_of_not_summable hf, + tsum_eq_zero_of_not_summable (summable_coe.not.mpr hf)] }, +end + +end infinite_sum + end quaternion From 13e18cfa070ea337ea960176414f5ae3a1534aae Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 17 Mar 2023 00:32:34 +0000 Subject: [PATCH 0644/1291] chore(algebra/ring/ulift): Split off and golf field instances (#18590) Move field-like instances on `ulift` from `algebra.ring.ulift` to a new file `algebra.field.ulift`. Golf them by declaring the `has_nat_cast` and `has_int_cast` instances earlier. --- src/algebra/field/ulift.lean | 39 ++++++++++++++++++++++++++++ src/algebra/group/ulift.lean | 42 ++++++++++++++++++------------- src/algebra/ring/ulift.lean | 22 ---------------- src/field_theory/cardinality.lean | 4 +-- 4 files changed, 65 insertions(+), 42 deletions(-) create mode 100644 src/algebra/field/ulift.lean diff --git a/src/algebra/field/ulift.lean b/src/algebra/field/ulift.lean new file mode 100644 index 0000000000000..fe49311689450 --- /dev/null +++ b/src/algebra/field/ulift.lean @@ -0,0 +1,39 @@ +/- +Copyright (c) 2023 Yaël Dillies. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies +-/ +import algebra.field.basic +import algebra.ring.ulift + +/-! +# Field instances for `ulift` + +This file defines instances for field, semifield and related structures on `ulift` types. + +(Recall `ulift α` is just a "copy" of a type `α` in a higher universe.) +-/ + +universes u v +variables {α : Type u} {x y : ulift.{v} α} + +namespace ulift + +instance [has_rat_cast α] : has_rat_cast (ulift α) := ⟨λ a, up a⟩ + +@[simp, norm_cast] lemma up_rat_cast [has_rat_cast α] (q : ℚ) : up (q : α) = q := rfl +@[simp, norm_cast] lemma down_rat_cast [has_rat_cast α] (q : ℚ) : down (q : ulift α) = q := rfl + +instance division_semiring [division_semiring α] : division_semiring (ulift α) := +by refine down_injective.division_semiring down _ _ _ _ _ _ _ _ _ _; intros; refl + +instance semifield [semifield α] : semifield (ulift α) := +{ ..ulift.division_semiring, ..ulift.comm_group_with_zero } + +instance division_ring [division_ring α] : division_ring (ulift α) := +{ ..ulift.division_semiring, ..ulift.add_group } + +instance field [field α] : field (ulift α) := +{ ..ulift.semifield, ..ulift.division_ring } + +end ulift diff --git a/src/algebra/group/ulift.lean b/src/algebra/group/ulift.lean index 49170b7c835cc..c90c9083e88f2 100644 --- a/src/algebra/group/ulift.lean +++ b/src/algebra/group/ulift.lean @@ -77,20 +77,27 @@ equiv.ulift.injective.mul_zero_one_class _ rfl rfl $ λ x y, rfl instance monoid [monoid α] : monoid (ulift α) := equiv.ulift.injective.monoid _ rfl (λ _ _, rfl) (λ _ _, rfl) -instance add_monoid_with_one [add_monoid_with_one α] : add_monoid_with_one (ulift α) := -{ nat_cast := λ n, ⟨n⟩, - nat_cast_zero := congr_arg ulift.up nat.cast_zero, - nat_cast_succ := λ n, congr_arg ulift.up (nat.cast_succ _), - .. ulift.has_one, .. ulift.add_monoid } - -@[simp] lemma nat_cast_down [add_monoid_with_one α] (n : ℕ) : - (n : ulift α).down = n := -rfl - @[to_additive] instance comm_monoid [comm_monoid α] : comm_monoid (ulift α) := equiv.ulift.injective.comm_monoid _ rfl (λ _ _, rfl) (λ _ _, rfl) +instance [has_nat_cast α] : has_nat_cast (ulift α) := ⟨λ n, up n⟩ +instance [has_int_cast α] : has_int_cast (ulift α) := ⟨λ n, up n⟩ + +@[simp, norm_cast] lemma up_nat_cast [has_nat_cast α] (n : ℕ) : up (n : α) = n := rfl +@[simp, norm_cast] lemma up_int_cast [has_int_cast α] (n : ℤ) : up (n : α) = n := rfl +@[simp, norm_cast] lemma down_nat_cast [has_nat_cast α] (n : ℕ) : down (n : ulift α) = n := rfl +@[simp, norm_cast] lemma down_int_cast [has_int_cast α] (n : ℤ) : down (n : ulift α) = n := rfl + +instance add_monoid_with_one [add_monoid_with_one α] : add_monoid_with_one (ulift α) := +{ nat_cast_zero := congr_arg ulift.up nat.cast_zero, + nat_cast_succ := λ n, congr_arg ulift.up (nat.cast_succ _), + .. ulift.has_one, .. ulift.add_monoid, ..ulift.has_nat_cast } + +instance add_comm_monoid_with_one [add_comm_monoid_with_one α] : + add_comm_monoid_with_one (ulift α) := +{ ..ulift.add_monoid_with_one, .. ulift.add_comm_monoid } + instance monoid_with_zero [monoid_with_zero α] : monoid_with_zero (ulift α) := equiv.ulift.injective.monoid_with_zero _ rfl rfl (λ _ _, rfl) (λ _ _, rfl) @@ -107,20 +114,19 @@ instance group [group α] : group (ulift α) := equiv.ulift.injective.group _ rfl (λ _ _, rfl) (λ _, rfl) (λ _ _, rfl) (λ _ _, rfl) (λ _ _, rfl) +@[to_additive] +instance comm_group [comm_group α] : comm_group (ulift α) := +equiv.ulift.injective.comm_group _ rfl (λ _ _, rfl) (λ _, rfl) + (λ _ _, rfl) (λ _ _, rfl) (λ _ _, rfl) + instance add_group_with_one [add_group_with_one α] : add_group_with_one (ulift α) := { int_cast := λ n, ⟨n⟩, int_cast_of_nat := λ n, congr_arg ulift.up (int.cast_of_nat _), int_cast_neg_succ_of_nat := λ n, congr_arg ulift.up (int.cast_neg_succ_of_nat _), .. ulift.add_monoid_with_one, .. ulift.add_group } -@[simp] lemma int_cast_down [add_group_with_one α] (n : ℤ) : - (n : ulift α).down = n := -rfl - -@[to_additive] -instance comm_group [comm_group α] : comm_group (ulift α) := -equiv.ulift.injective.comm_group _ rfl (λ _ _, rfl) (λ _, rfl) - (λ _ _, rfl) (λ _ _, rfl) (λ _ _, rfl) +instance add_comm_group_with_one [add_comm_group_with_one α] : add_comm_group_with_one (ulift α) := +{ ..ulift.add_group_with_one, .. ulift.add_comm_group } instance group_with_zero [group_with_zero α] : group_with_zero (ulift α) := equiv.ulift.injective.group_with_zero _ rfl rfl (λ _ _, rfl) (λ _, rfl) (λ _ _, rfl) (λ _ _, rfl) diff --git a/src/algebra/ring/ulift.lean b/src/algebra/ring/ulift.lean index 6336fa3f6fe76..ea974bbd8c39e 100644 --- a/src/algebra/ring/ulift.lean +++ b/src/algebra/ring/ulift.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ import algebra.group.ulift -import algebra.field.defs import algebra.ring.equiv /-! @@ -107,25 +106,4 @@ instance comm_ring [comm_ring α] : comm_ring (ulift α) := by refine_struct { .. ulift.ring }; tactic.pi_instance_derive_field -instance [has_rat_cast α] : has_rat_cast (ulift α) := -⟨λ a, ulift.up (coe a)⟩ - -@[simp] lemma rat_cast_down [has_rat_cast α] (n : ℚ) : ulift.down (n : ulift α) = n := -rfl - -instance field [field α] : field (ulift α) := -begin - have of_rat_mk : ∀ a b h1 h2, ((⟨a, b, h1, h2⟩ : ℚ) : ulift α) = ↑a * (↑b)⁻¹, - { intros a b h1 h2, - ext, - rw [rat_cast_down, mul_down, inv_down, nat_cast_down, int_cast_down], - exact field.rat_cast_mk a b h1 h2 }, - refine_struct { zero := (0 : ulift α), inv := has_inv.inv, div := has_div.div, - zpow := λ n a, ulift.up (a.down ^ n), rat_cast := coe, rat_cast_mk := of_rat_mk, qsmul := (•), - .. @ulift.nontrivial α _, .. ulift.comm_ring }; tactic.pi_instance_derive_field, - -- `mul_inv_cancel` requires special attention: it leaves the goal `∀ {a}, a ≠ 0 → a * a⁻¹ = 1`. - cases a, - tauto -end - end ulift diff --git a/src/field_theory/cardinality.lean b/src/field_theory/cardinality.lean index 049050e06e135..9fe35f450678a 100644 --- a/src/field_theory/cardinality.lean +++ b/src/field_theory/cardinality.lean @@ -3,14 +3,14 @@ Copyright (c) 2022 Eric Rodriguez. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Eric Rodriguez -/ -import algebra.ring.ulift +import algebra.field.ulift import data.mv_polynomial.cardinal +import data.nat.factorization.prime_pow import data.rat.denumerable import field_theory.finite.galois_field import logic.equiv.transfer_instance import ring_theory.localization.cardinality import set_theory.cardinal.divisibility -import data.nat.factorization.prime_pow /-! # Cardinality of Fields From 30413fc89f202a090a54d78e540963ed3de0056e Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 17 Mar 2023 09:39:09 +0000 Subject: [PATCH 0645/1291] chore(algebra): generalize typeclass arguments from field to semifield (#18597) This generalizes some typeclass arguments from `field` to `semifield` and `division_ring` to `division_semiring`. The proof for `map_inv_nat_cast_smul` had to be rewritten, as it was previously proved in terms of `map_inv_int_cast_smul`. The latter is now instead proved in terms of the former. Forward-ported in https://github.com/leanprover-community/mathlib4/pull/2926 Co-authored-by: Eric Wieser --- src/algebra/module/basic.lean | 53 ++++++++++++++++-------------- src/algebra/periodic.lean | 33 ++++++++++--------- src/algebra/star/basic.lean | 6 ++-- src/algebra/star/module.lean | 14 ++++---- src/algebra/star/pointwise.lean | 2 +- src/algebra/star/self_adjoint.lean | 9 +++-- src/data/matrix/basic.lean | 2 +- 7 files changed, 65 insertions(+), 54 deletions(-) diff --git a/src/algebra/module/basic.lean b/src/algebra/module/basic.lean index 14d70984fe3a3..243f4ca3953a4 100644 --- a/src/algebra/module/basic.lean +++ b/src/algebra/module/basic.lean @@ -381,27 +381,32 @@ lemma map_nat_cast_smul [add_comm_monoid M] [add_comm_monoid M₂] {F : Type*} f ((x : R) • a) = (x : S) • f a := by simp only [←nsmul_eq_smul_cast, map_nsmul] -lemma map_inv_int_cast_smul [add_comm_group M] [add_comm_group M₂] {F : Type*} +lemma map_inv_nat_cast_smul [add_comm_monoid M] [add_comm_monoid M₂] {F : Type*} [add_monoid_hom_class F M M₂] (f : F) - (R S : Type*) [division_ring R] [division_ring S] [module R M] [module S M₂] - (n : ℤ) (x : M) : + (R S : Type*) [division_semiring R] [division_semiring S] [module R M] [module S M₂] + (n : ℕ) (x : M) : f ((n⁻¹ : R) • x) = (n⁻¹ : S) • f x := begin by_cases hR : (n : R) = 0; by_cases hS : (n : S) = 0, { simp [hR, hS] }, { suffices : ∀ y, f y = 0, by simp [this], clear x, intro x, - rw [← inv_smul_smul₀ hS (f x), ← map_int_cast_smul f R S], simp [hR] }, + rw [← inv_smul_smul₀ hS (f x), ← map_nat_cast_smul f R S], simp [hR] }, { suffices : ∀ y, f y = 0, by simp [this], clear x, intro x, - rw [← smul_inv_smul₀ hR x, map_int_cast_smul f R S, hS, zero_smul] }, - { rw [← inv_smul_smul₀ hS (f _), ← map_int_cast_smul f R S, smul_inv_smul₀ hR] } + rw [← smul_inv_smul₀ hR x, map_nat_cast_smul f R S, hS, zero_smul] }, + { rw [← inv_smul_smul₀ hS (f _), ← map_nat_cast_smul f R S, smul_inv_smul₀ hR] } end -lemma map_inv_nat_cast_smul [add_comm_group M] [add_comm_group M₂] {F : Type*} +lemma map_inv_int_cast_smul [add_comm_group M] [add_comm_group M₂] {F : Type*} [add_monoid_hom_class F M M₂] (f : F) (R S : Type*) [division_ring R] [division_ring S] [module R M] [module S M₂] - (n : ℕ) (x : M) : - f ((n⁻¹ : R) • x) = (n⁻¹ : S) • f x := -by exact_mod_cast map_inv_int_cast_smul f R S n x + (z : ℤ) (x : M) : + f ((z⁻¹ : R) • x) = (z⁻¹ : S) • f x := +begin + obtain ⟨n, rfl | rfl⟩ := z.eq_coe_or_neg, + { rw [int.cast_coe_nat, int.cast_coe_nat, map_inv_nat_cast_smul _ R S] }, + { simp_rw [int.cast_neg, int.cast_coe_nat, inv_neg, neg_smul, map_neg, + map_inv_nat_cast_smul _ R S] }, +end lemma map_rat_cast_smul [add_comm_group M] [add_comm_group M₂] {F : Type*} [add_monoid_hom_class F M M₂] (f : F) @@ -421,6 +426,13 @@ instance subsingleton_rat_module (E : Type*) [add_comm_group E] : subsingleton ( ⟨λ P Q, module.ext' P Q $ λ r x, @map_rat_smul _ _ _ _ P Q _ _ (add_monoid_hom.id E) r x⟩ +/-- If `E` is a vector space over two division semirings `R` and `S`, then scalar multiplications +agree on inverses of natural numbers in `R` and `S`. -/ +lemma inv_nat_cast_smul_eq {E : Type*} (R S : Type*) [add_comm_monoid E] [division_semiring R] + [division_semiring S] [module R E] [module S E] (n : ℕ) (x : E) : + (n⁻¹ : R) • x = (n⁻¹ : S) • x := +map_inv_nat_cast_smul (add_monoid_hom.id E) R S n x + /-- If `E` is a vector space over two division rings `R` and `S`, then scalar multiplications agree on inverses of integer numbers in `R` and `S`. -/ lemma inv_int_cast_smul_eq {E : Type*} (R S : Type*) [add_comm_group E] [division_ring R] @@ -428,27 +440,20 @@ lemma inv_int_cast_smul_eq {E : Type*} (R S : Type*) [add_comm_group E] [divisio (n⁻¹ : R) • x = (n⁻¹ : S) • x := map_inv_int_cast_smul (add_monoid_hom.id E) R S n x -/-- If `E` is a vector space over two division rings `R` and `S`, then scalar multiplications -agree on inverses of natural numbers in `R` and `S`. -/ -lemma inv_nat_cast_smul_eq {E : Type*} (R S : Type*) [add_comm_group E] [division_ring R] - [division_ring S] [module R E] [module S E] (n : ℕ) (x : E) : - (n⁻¹ : R) • x = (n⁻¹ : S) • x := -map_inv_nat_cast_smul (add_monoid_hom.id E) R S n x +/-- If `E` is a vector space over a division ring `R` and has a monoid action by `α`, then that +action commutes by scalar multiplication of inverses of natural numbers in `R`. -/ +lemma inv_nat_cast_smul_comm {α E : Type*} (R : Type*) [add_comm_monoid E] [division_semiring R] + [monoid α] [module R E] [distrib_mul_action α E] (n : ℕ) (s : α) (x : E) : + (n⁻¹ : R) • s • x = s • (n⁻¹ : R) • x := +(map_inv_nat_cast_smul (distrib_mul_action.to_add_monoid_hom E s) R R n x).symm -/-- If `E` is a vector space over a division rings `R` and has a monoid action by `α`, then that +/-- If `E` is a vector space over a division ring `R` and has a monoid action by `α`, then that action commutes by scalar multiplication of inverses of integers in `R` -/ lemma inv_int_cast_smul_comm {α E : Type*} (R : Type*) [add_comm_group E] [division_ring R] [monoid α] [module R E] [distrib_mul_action α E] (n : ℤ) (s : α) (x : E) : (n⁻¹ : R) • s • x = s • (n⁻¹ : R) • x := (map_inv_int_cast_smul (distrib_mul_action.to_add_monoid_hom E s) R R n x).symm -/-- If `E` is a vector space over a division rings `R` and has a monoid action by `α`, then that -action commutes by scalar multiplication of inverses of natural numbers in `R`. -/ -lemma inv_nat_cast_smul_comm {α E : Type*} (R : Type*) [add_comm_group E] [division_ring R] - [monoid α] [module R E] [distrib_mul_action α E] (n : ℕ) (s : α) (x : E) : - (n⁻¹ : R) • s • x = s • (n⁻¹ : R) • x := -(map_inv_nat_cast_smul (distrib_mul_action.to_add_monoid_hom E s) R R n x).symm - /-- If `E` is a vector space over two division rings `R` and `S`, then scalar multiplications agree on rational numbers in `R` and `S`. -/ lemma rat_cast_smul_eq {E : Type*} (R S : Type*) [add_comm_group E] [division_ring R] diff --git a/src/algebra/periodic.lean b/src/algebra/periodic.lean index d5bd16c586aae..1ca286569804a 100644 --- a/src/algebra/periodic.lean +++ b/src/algebra/periodic.lean @@ -103,7 +103,7 @@ protected lemma periodic.const_smul [add_monoid α] [group γ] [distrib_mul_acti periodic (λ x, f (a • x)) (a⁻¹ • c) := λ x, by simpa only [smul_add, smul_inv_smul] using h (a • x) -lemma periodic.const_smul₀ [add_comm_monoid α] [division_ring γ] [module γ α] +lemma periodic.const_smul₀ [add_comm_monoid α] [division_semiring γ] [module γ α] (h : periodic f c) (a : γ) : periodic (λ x, f (a • x)) (a⁻¹ • c) := begin @@ -112,7 +112,7 @@ begin simpa only [smul_add, smul_inv_smul₀ ha] using h (a • x), end -protected lemma periodic.const_mul [division_ring α] (h : periodic f c) (a : α) : +protected lemma periodic.const_mul [division_semiring α] (h : periodic f c) (a : α) : periodic (λ x, f (a * x)) (a⁻¹ * c) := h.const_smul₀ a @@ -121,29 +121,29 @@ lemma periodic.const_inv_smul [add_monoid α] [group γ] [distrib_mul_action γ periodic (λ x, f (a⁻¹ • x)) (a • c) := by simpa only [inv_inv] using h.const_smul a⁻¹ -lemma periodic.const_inv_smul₀ [add_comm_monoid α] [division_ring γ] [module γ α] +lemma periodic.const_inv_smul₀ [add_comm_monoid α] [division_semiring γ] [module γ α] (h : periodic f c) (a : γ) : periodic (λ x, f (a⁻¹ • x)) (a • c) := by simpa only [inv_inv] using h.const_smul₀ a⁻¹ -lemma periodic.const_inv_mul [division_ring α] (h : periodic f c) (a : α) : +lemma periodic.const_inv_mul [division_semiring α] (h : periodic f c) (a : α) : periodic (λ x, f (a⁻¹ * x)) (a * c) := h.const_inv_smul₀ a -lemma periodic.mul_const [division_ring α] (h : periodic f c) (a : α) : +lemma periodic.mul_const [division_semiring α] (h : periodic f c) (a : α) : periodic (λ x, f (x * a)) (c * a⁻¹) := h.const_smul₀ $ mul_opposite.op a -lemma periodic.mul_const' [division_ring α] +lemma periodic.mul_const' [division_semiring α] (h : periodic f c) (a : α) : periodic (λ x, f (x * a)) (c / a) := by simpa only [div_eq_mul_inv] using h.mul_const a -lemma periodic.mul_const_inv [division_ring α] (h : periodic f c) (a : α) : +lemma periodic.mul_const_inv [division_semiring α] (h : periodic f c) (a : α) : periodic (λ x, f (x * a⁻¹)) (c * a) := h.const_inv_smul₀ $ mul_opposite.op a -lemma periodic.div_const [division_ring α] (h : periodic f c) (a : α) : +lemma periodic.div_const [division_semiring α] (h : periodic f c) (a : α) : periodic (λ x, f (x / a)) (c * a) := by simpa only [div_eq_mul_inv] using h.mul_const_inv a @@ -425,12 +425,12 @@ lemma antiperiodic.const_smul [add_monoid α] [has_neg β] [group γ] [distrib_m antiperiodic (λ x, f (a • x)) (a⁻¹ • c) := λ x, by simpa only [smul_add, smul_inv_smul] using h (a • x) -lemma antiperiodic.const_smul₀ [add_comm_monoid α] [has_neg β] [division_ring γ] [module γ α] +lemma antiperiodic.const_smul₀ [add_comm_monoid α] [has_neg β] [division_semiring γ] [module γ α] (h : antiperiodic f c) {a : γ} (ha : a ≠ 0) : antiperiodic (λ x, f (a • x)) (a⁻¹ • c) := λ x, by simpa only [smul_add, smul_inv_smul₀ ha] using h (a • x) -lemma antiperiodic.const_mul [division_ring α] [has_neg β] +lemma antiperiodic.const_mul [division_semiring α] [has_neg β] (h : antiperiodic f c) {a : α} (ha : a ≠ 0) : antiperiodic (λ x, f (a * x)) (a⁻¹ * c) := h.const_smul₀ ha @@ -440,32 +440,33 @@ lemma antiperiodic.const_inv_smul [add_monoid α] [has_neg β] [group γ] [distr antiperiodic (λ x, f (a⁻¹ • x)) (a • c) := by simpa only [inv_inv] using h.const_smul a⁻¹ -lemma antiperiodic.const_inv_smul₀ [add_comm_monoid α] [has_neg β] [division_ring γ] [module γ α] +lemma antiperiodic.const_inv_smul₀ + [add_comm_monoid α] [has_neg β] [division_semiring γ] [module γ α] (h : antiperiodic f c) {a : γ} (ha : a ≠ 0) : antiperiodic (λ x, f (a⁻¹ • x)) (a • c) := by simpa only [inv_inv] using h.const_smul₀ (inv_ne_zero ha) -lemma antiperiodic.const_inv_mul [division_ring α] [has_neg β] +lemma antiperiodic.const_inv_mul [division_semiring α] [has_neg β] (h : antiperiodic f c) {a : α} (ha : a ≠ 0) : antiperiodic (λ x, f (a⁻¹ * x)) (a * c) := h.const_inv_smul₀ ha -lemma antiperiodic.mul_const [division_ring α] [has_neg β] +lemma antiperiodic.mul_const [division_semiring α] [has_neg β] (h : antiperiodic f c) {a : α} (ha : a ≠ 0) : antiperiodic (λ x, f (x * a)) (c * a⁻¹) := h.const_smul₀ $ (mul_opposite.op_ne_zero_iff a).mpr ha -lemma antiperiodic.mul_const' [division_ring α] [has_neg β] +lemma antiperiodic.mul_const' [division_semiring α] [has_neg β] (h : antiperiodic f c) {a : α} (ha : a ≠ 0) : antiperiodic (λ x, f (x * a)) (c / a) := by simpa only [div_eq_mul_inv] using h.mul_const ha -lemma antiperiodic.mul_const_inv [division_ring α] [has_neg β] +lemma antiperiodic.mul_const_inv [division_semiring α] [has_neg β] (h : antiperiodic f c) {a : α} (ha : a ≠ 0) : antiperiodic (λ x, f (x * a⁻¹)) (c * a) := h.const_inv_smul₀ $ (mul_opposite.op_ne_zero_iff a).mpr ha -protected lemma antiperiodic.div_inv [division_ring α] [has_neg β] +protected lemma antiperiodic.div_inv [division_semiring α] [has_neg β] (h : antiperiodic f c) {a : α} (ha : a ≠ 0) : antiperiodic (λ x, f (x / a)) (c * a) := by simpa only [div_eq_mul_inv] using h.mul_const_inv ha diff --git a/src/algebra/star/basic.lean b/src/algebra/star/basic.lean index 6b74729fcc6b4..de34a4d606ac6 100644 --- a/src/algebra/star/basic.lean +++ b/src/algebra/star/basic.lean @@ -310,15 +310,15 @@ lemma ring_hom.star_apply {S : Type*} [non_assoc_semiring S] [comm_semiring R] [ alias star_ring_end_self_apply ← complex.conj_conj alias star_ring_end_self_apply ← is_R_or_C.conj_conj -@[simp] lemma star_inv' [division_ring R] [star_ring R] (x : R) : star (x⁻¹) = (star x)⁻¹ := +@[simp] lemma star_inv' [division_semiring R] [star_ring R] (x : R) : star (x⁻¹) = (star x)⁻¹ := op_injective $ (map_inv₀ (star_ring_equiv : R ≃+* Rᵐᵒᵖ) x).trans (op_inv (star x)).symm -@[simp] lemma star_zpow₀ [division_ring R] [star_ring R] (x : R) (z : ℤ) : +@[simp] lemma star_zpow₀ [division_semiring R] [star_ring R] (x : R) (z : ℤ) : star (x ^ z) = star x ^ z := op_injective $ (map_zpow₀ (star_ring_equiv : R ≃+* Rᵐᵒᵖ) x z).trans (op_zpow (star x) z).symm /-- When multiplication is commutative, `star` preserves division. -/ -@[simp] lemma star_div' [field R] [star_ring R] (x y : R) : star (x / y) = star x / star y := +@[simp] lemma star_div' [semifield R] [star_ring R] (x y : R) : star (x / y) = star x / star y := map_div₀ (star_ring_end R) _ _ @[simp] lemma star_bit0 [add_monoid R] [star_add_monoid R] (r : R) : diff --git a/src/algebra/star/module.lean b/src/algebra/star/module.lean index 321ea313b9498..17e971650b8e8 100644 --- a/src/algebra/star/module.lean +++ b/src/algebra/star/module.lean @@ -33,22 +33,22 @@ This file also provides some lemmas that need `algebra.module.basic` imported to section smul_lemmas variables {R M : Type*} +@[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_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_nat_cast_smul [division_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_inv_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 diff --git a/src/algebra/star/pointwise.lean b/src/algebra/star/pointwise.lean index f4b883d973333..4885c23b117cb 100644 --- a/src/algebra/star/pointwise.lean +++ b/src/algebra/star/pointwise.lean @@ -117,7 +117,7 @@ instance [has_star α] [has_trivial_star α] : has_trivial_star (set α) := protected lemma star_inv [group α] [star_semigroup α] (s : set α) : (s⁻¹)⋆ = (s⋆)⁻¹ := by { ext, simp only [mem_star, mem_inv, star_inv] } -protected lemma star_inv' [division_ring α] [star_ring α] (s : set α) : (s⁻¹)⋆ = (s⋆)⁻¹ := +protected lemma star_inv' [division_semiring α] [star_ring α] (s : set α) : (s⁻¹)⋆ = (s⋆)⁻¹ := by { ext, simp only [mem_star, mem_inv, star_inv'] } end set diff --git a/src/algebra/star/self_adjoint.lean b/src/algebra/star/self_adjoint.lean index d0dc47e4e4988..47d431e7355dc 100644 --- a/src/algebra/star/self_adjoint.lean +++ b/src/algebra/star/self_adjoint.lean @@ -166,8 +166,8 @@ star_int_cast _ end ring -section division_ring -variables [division_ring R] [star_ring R] +section division_semiring +variables [division_semiring R] [star_ring R] lemma inv {x : R} (hx : is_self_adjoint x) : is_self_adjoint x⁻¹ := by simp only [is_self_adjoint_iff, star_inv', hx.star_eq] @@ -175,6 +175,11 @@ by simp only [is_self_adjoint_iff, star_inv', hx.star_eq] lemma zpow {x : R} (hx : is_self_adjoint x) (n : ℤ) : is_self_adjoint (x ^ n):= by simp only [is_self_adjoint_iff, star_zpow₀, hx.star_eq] +end division_semiring + +section division_ring +variables [division_ring R] [star_ring R] + lemma _root_.is_self_adjoint_rat_cast (x : ℚ) : is_self_adjoint (x : R) := star_rat_cast _ diff --git a/src/data/matrix/basic.lean b/src/data/matrix/basic.lean index ed5963228f606..85e12b44c0ac0 100644 --- a/src/data/matrix/basic.lean +++ b/src/data/matrix/basic.lean @@ -1537,7 +1537,7 @@ matrix.ext $ by simp [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 α] +@[simp] lemma conj_transpose_inv_nat_cast_smul [division_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 From 2d5739b61641ee4e7e53eca5688a08f66f2e6a60 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 17 Mar 2023 12:30:23 +0000 Subject: [PATCH 0646/1291] chore(ring_theory/power_series/basic): remove commutativity assumption (#18599) Also moves a lost lemma that belongs in an earlier file. --- src/data/mv_polynomial/basic.lean | 3 +++ src/ring_theory/power_series/basic.lean | 19 +++++++++---------- 2 files changed, 12 insertions(+), 10 deletions(-) diff --git a/src/data/mv_polynomial/basic.lean b/src/data/mv_polynomial/basic.lean index a361af8d54c34..4d4fefcd20717 100644 --- a/src/data/mv_polynomial/basic.lean +++ b/src/data/mv_polynomial/basic.lean @@ -1099,6 +1099,9 @@ section aeval variables [algebra R S₁] [comm_semiring S₂] variables (f : σ → S₁) +lemma algebra_map_apply (r : R) : + algebra_map R (mv_polynomial σ S₁) r = C (algebra_map R S₁ r) := rfl + /-- A map `σ → S₁` where `S₁` is an algebra over `R` generates an `R`-algebra homomorphism from multivariate polynomials over `σ` to `S₁`. -/ def aeval : mv_polynomial σ R →ₐ[R] S₁ := diff --git a/src/ring_theory/power_series/basic.lean b/src/ring_theory/power_series/basic.lean index 5752dc54e58b4..591e43568d3f0 100644 --- a/src/ring_theory/power_series/basic.lean +++ b/src/ring_theory/power_series/basic.lean @@ -542,8 +542,8 @@ end end trunc -section comm_semiring -variable [comm_semiring R] +section semiring +variable [semiring R] lemma X_pow_dvd_iff {s : σ} {n : ℕ} {φ : mv_power_series σ R} : (X s : mv_power_series σ R)^n ∣ φ ↔ ∀ m : σ →₀ ℕ, m s < n → coeff R m φ = 0 := @@ -585,7 +585,8 @@ begin { exact h m (hm.symm ▸ zero_lt_one) }, { exact h m (nat.eq_zero_of_le_zero $ nat.le_of_succ_le_succ hm) } end -end comm_semiring + +end semiring section ring variables [ring R] @@ -886,8 +887,6 @@ section algebra variables (A : Type*) [comm_semiring A] [algebra R A] -lemma algebra_map_apply (r : R) : algebra_map R (mv_polynomial σ A) r = C (algebra_map R A r) := rfl - /-- The coercion from multivariable polynomials to multivariable power series as an algebra homomorphism. @@ -1241,11 +1240,6 @@ by { ext, simp [coeff_X, apply_ite f] } end map -end semiring - -section comm_semiring -variables [comm_semiring R] - lemma X_pow_dvd_iff {n : ℕ} {φ : power_series R} : (X : power_series R)^n ∣ φ ↔ ∀ m, m < n → coeff R m φ = 0 := begin @@ -1264,6 +1258,11 @@ begin { intros m hm, rwa nat.eq_zero_of_le_zero (nat.le_of_succ_le_succ hm) } end +end semiring + +section comm_semiring +variables [comm_semiring R] + open finset nat /-- The ring homomorphism taking a power series `f(X)` to `f(aX)`. -/ From 759575657f189ccb424b990164c8b1fa9f55cdfe Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 17 Mar 2023 12:30:25 +0000 Subject: [PATCH 0647/1291] chore(data/fintype/fin): add Iio lemmas (#18600) These are the analogue to the existing `Ioi` lemmas in this file. `Iio_last_eq_map` is the dual of `Ioi_zero_eq_map`, and `Iio_cast_succ` is the dual of `Ioi_succ`. I couldn't find a better home for `map_subtype_embedding_univ`. Note that it is deliberately `subtype_embedding` and not `coe_embedding` to match the similarly-misnamed `fin.map_subtype_embedding_Iio` --- src/data/fintype/fin.lean | 24 ++++++++++++++++++++++++ 1 file changed, 24 insertions(+) diff --git a/src/data/fintype/fin.lean b/src/data/fintype/fin.lean index 7c834516819b4..1a68e3bf35ae7 100644 --- a/src/data/fintype/fin.lean +++ b/src/data/fintype/fin.lean @@ -23,6 +23,14 @@ namespace fin variables {α β : Type*} {n : ℕ} +-- TODO: replace `subtype` with `coe` in the name of this lemma and `fin.map_subtype_embedding_Iio` +lemma map_subtype_embedding_univ : + (finset.univ : finset (fin n)).map fin.coe_embedding = Iio n := +begin + ext, + simp [order_iso_subtype.symm.surjective.exists, order_iso.symm], +end + @[simp] lemma Ioi_zero_eq_map : Ioi (0 : fin n.succ) = univ.map (fin.succ_embedding _).to_embedding := begin @@ -36,6 +44,14 @@ begin exact succ_pos _ }, end +@[simp] lemma Iio_last_eq_map : + Iio (fin.last n) = finset.univ.map fin.cast_succ.to_embedding := +begin + apply finset.map_injective fin.coe_embedding, + rw [finset.map_map, fin.map_subtype_embedding_Iio, fin.coe_last], + exact map_subtype_embedding_univ.symm +end + @[simp] lemma Ioi_succ (i : fin n) : Ioi i.succ = (Ioi i).map (fin.succ_embedding _).to_embedding := begin @@ -50,6 +66,14 @@ begin { rintro ⟨i, hi, rfl⟩, simpa }, end +@[simp] lemma Iio_cast_succ (i : fin n) : + Iio (cast_succ i) = (Iio i).map fin.cast_succ.to_embedding := +begin + apply finset.map_injective fin.coe_embedding, + rw [finset.map_map, fin.map_subtype_embedding_Iio], + exact (fin.map_subtype_embedding_Iio i).symm, +end + lemma card_filter_univ_succ' (p : fin (n + 1) → Prop) [decidable_pred p] : (univ.filter p).card = (ite (p 0) 1 0) + (univ.filter (p ∘ fin.succ)).card := begin From 88b8a77d63a702923d9bee05e9e454ebc22aa766 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 17 Mar 2023 14:58:09 +0000 Subject: [PATCH 0648/1291] chore(topology/basic): backport a generalization to Sort (#18544) mathport currently complains that these don't align. --- src/topology/basic.lean | 5 ++--- 1 file changed, 2 insertions(+), 3 deletions(-) diff --git a/src/topology/basic.lean b/src/topology/basic.lean index c704a3fd2f161..da690183da981 100644 --- a/src/topology/basic.lean +++ b/src/topology/basic.lean @@ -143,9 +143,8 @@ finite.induction_on hs (λ a s has hs ih h, by rw bInter_insert; exact is_open.inter (h a (mem_insert _ _)) (ih (λ i hi, h i (mem_insert_of_mem _ hi)))) -lemma is_open_Inter [finite β] {s : β → set α} (h : ∀ i, is_open (s i)) : is_open (⋂ i, s i) := -suffices is_open (⋂ (i : β) (hi : i ∈ @univ β), s i), by simpa, -is_open_bInter finite_univ (λ i _, h i) +lemma is_open_Inter [finite ι] {s : ι → set α} (h : ∀ i, is_open (s i)) : is_open (⋂ i, s i) := +is_open_sInter (finite_range _) (forall_range_iff.2 h) lemma is_open_Inter_prop {p : Prop} {s : p → set α} (h : ∀ h : p, is_open (s h)) : is_open (Inter s) := From 9f0d61b4475e3c3cba6636ab51cdb1f3949d2e1d Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 17 Mar 2023 14:58:10 +0000 Subject: [PATCH 0649/1291] fix(analysis/inner_product_space/pi_L2): add missing type cast functions (#18574) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The preferred way to convert between `ι → 𝕜` and `euclidean_space 𝕜 ι` is via `pi_Lp.equiv`, as this preserves the right typing information. We use the local notation `⟪x, y⟫ₑ` to refer the euclidean inner product of `x y : ι → 𝕜`, which inserts the casting within the notation. This adds a new definition `matrix.to_euclidean_lin` as a shorthand to turn a matrix into a `linear_map` over `euclidean_space`. It also generalizes `inner_matrix_row_row` and `inner_matrix_col_col` away from `fin n` to arbitrary index types. --- src/analysis/inner_product_space/adjoint.lean | 11 ++--- src/analysis/inner_product_space/pi_L2.lean | 46 +++++++++++++++---- src/linear_algebra/matrix/hermitian.lean | 11 ++--- src/linear_algebra/matrix/ldl.lean | 15 +++--- src/linear_algebra/matrix/spectrum.lean | 16 +++---- 5 files changed, 58 insertions(+), 41 deletions(-) diff --git a/src/analysis/inner_product_space/adjoint.lean b/src/analysis/inner_product_space/adjoint.lean index f5a26f01b4cbc..12ffdcc25e197 100644 --- a/src/analysis/inner_product_space/adjoint.lean +++ b/src/analysis/inner_product_space/adjoint.lean @@ -434,14 +434,13 @@ open_locale complex_conjugate /-- The adjoint of the linear map associated to a matrix is the linear map associated to the conjugate transpose of that matrix. -/ -lemma conj_transpose_eq_adjoint (A : matrix m n 𝕜) : - to_lin' A.conj_transpose = - @linear_map.adjoint _ (euclidean_space 𝕜 n) (euclidean_space 𝕜 m) _ _ _ _ _ (to_lin' A) := +lemma to_euclidean_lin_conj_transpose_eq_adjoint (A : matrix m n 𝕜) : + A.conj_transpose.to_euclidean_lin = A.to_euclidean_lin.adjoint := begin - rw @linear_map.eq_adjoint_iff _ (euclidean_space 𝕜 m) (euclidean_space 𝕜 n), + rw linear_map.eq_adjoint_iff, intros x y, - convert dot_product_assoc (conj ∘ (id x : m → 𝕜)) y A using 1, - simp [dot_product, mul_vec, ring_hom.map_sum, ← star_ring_end_apply, mul_comm], + simp_rw [euclidean_space.inner_eq_star_dot_product, pi_Lp_equiv_to_euclidean_lin, + to_lin'_apply, star_mul_vec, conj_transpose_conj_transpose, dot_product_mul_vec], end end matrix diff --git a/src/analysis/inner_product_space/pi_L2.lean b/src/analysis/inner_product_space/pi_L2.lean index 3b851881f4136..d294699838b12 100644 --- a/src/analysis/inner_product_space/pi_L2.lean +++ b/src/analysis/inner_product_space/pi_L2.lean @@ -138,6 +138,9 @@ lemma finrank_euclidean_space_fin {n : ℕ} : lemma euclidean_space.inner_eq_star_dot_product (x y : euclidean_space 𝕜 ι) : ⟪x, y⟫ = matrix.dot_product (star $ pi_Lp.equiv _ _ x) (pi_Lp.equiv _ _ y) := rfl +lemma euclidean_space.inner_pi_Lp_equiv_symm (x y : ι → 𝕜) : + ⟪(pi_Lp.equiv 2 _).symm x, (pi_Lp.equiv 2 _).symm y⟫ = matrix.dot_product (star x) y := rfl + /-- A finite, mutually orthogonal family of subspaces of `E`, which span `E`, induce an isometry from `E` to `pi_Lp 2` of the subspaces equipped with the `L2` inner product. -/ def direct_sum.is_internal.isometry_L2_of_orthogonal_family @@ -815,18 +818,41 @@ section matrix open_locale matrix -variables {n m : ℕ} +variables {m n : Type*} + +namespace matrix +variables [fintype m] [fintype n] [decidable_eq n] + +/-- `matrix.to_lin'` adapted for `euclidean_space 𝕜 _`. -/ +def to_euclidean_lin : matrix m n 𝕜 ≃ₗ[𝕜] (euclidean_space 𝕜 n →ₗ[𝕜] euclidean_space 𝕜 m) := +matrix.to_lin' ≪≫ₗ linear_equiv.arrow_congr + (pi_Lp.linear_equiv _ 𝕜 (λ _ : n, 𝕜)).symm (pi_Lp.linear_equiv _ 𝕜 (λ _ : m, 𝕜)).symm + +@[simp] +lemma to_euclidean_lin_pi_Lp_equiv_symm (A : matrix m n 𝕜) (x : n → 𝕜) : + A.to_euclidean_lin ((pi_Lp.equiv _ _).symm x) = (pi_Lp.equiv _ _).symm (A.to_lin' x) := rfl + +@[simp] +lemma pi_Lp_equiv_to_euclidean_lin (A : matrix m n 𝕜) (x : euclidean_space 𝕜 n) : + pi_Lp.equiv _ _ (A.to_euclidean_lin x) = A.to_lin' (pi_Lp.equiv _ _ x) := rfl + +/- `matrix.to_euclidean_lin` is the same as `matrix.to_lin` applied to `pi_Lp.basis_fun`, -/ +lemma to_euclidean_lin_eq_to_lin : + (to_euclidean_lin : matrix m n 𝕜 ≃ₗ[𝕜] _) = + matrix.to_lin (pi_Lp.basis_fun _ _ _) (pi_Lp.basis_fun _ _ _) := rfl + +end matrix -local notation `⟪`x`, `y`⟫ₘ` := @inner 𝕜 (euclidean_space 𝕜 (fin m)) _ x y -local notation `⟪`x`, `y`⟫ₙ` := @inner 𝕜 (euclidean_space 𝕜 (fin n)) _ x y +local notation `⟪`x`, `y`⟫ₑ` := @inner 𝕜 _ _ ((pi_Lp.equiv 2 _).symm x) ((pi_Lp.equiv 2 _).symm y) -/-- The inner product of a row of A and a row of B is an entry of B ⬝ Aᴴ. -/ -lemma inner_matrix_row_row (A B : matrix (fin n) (fin m) 𝕜) (i j : (fin n)) : - ⟪A i, B j⟫ₘ = (B ⬝ Aᴴ) j i := by {simp only [inner, matrix.mul_apply, star_ring_end_apply, - matrix.conj_transpose_apply,mul_comm]} +/-- The inner product of a row of `A` and a row of `B` is an entry of `B ⬝ Aᴴ`. -/ +lemma inner_matrix_row_row [fintype n] (A B : matrix m n 𝕜) (i j : m) : + ⟪A i, B j⟫ₑ = (B ⬝ Aᴴ) j i := +by simp_rw [euclidean_space.inner_pi_Lp_equiv_symm, matrix.mul_apply', matrix.dot_product_comm, + matrix.conj_transpose_apply, pi.star_def] -/-- The inner product of a column of A and a column of B is an entry of Aᴴ ⬝ B -/ -lemma inner_matrix_col_col (A B : matrix (fin n) (fin m) 𝕜) (i j : (fin m)) : - ⟪Aᵀ i, Bᵀ j⟫ₙ = (Aᴴ ⬝ B) i j := rfl +/-- The inner product of a column of `A` and a column of `B` is an entry of `Aᴴ ⬝ B`. -/ +lemma inner_matrix_col_col [fintype m] (A B : matrix m n 𝕜) (i j : n) : + ⟪Aᵀ i, Bᵀ j⟫ₑ = (Aᴴ ⬝ B) i j := rfl end matrix diff --git a/src/linear_algebra/matrix/hermitian.lean b/src/linear_algebra/matrix/hermitian.lean index d2d1767b917d4..4dace90d7d553 100644 --- a/src/linear_algebra/matrix/hermitian.lean +++ b/src/linear_algebra/matrix/hermitian.lean @@ -25,7 +25,7 @@ variables {α β : Type*} {m n : Type*} {A : matrix n n α} open_locale matrix -local notation `⟪`x`, `y`⟫` := @inner α (pi_Lp 2 (λ (_ : n), α)) _ x y +local notation `⟪`x`, `y`⟫` := @inner α _ _ x y section non_unital_semiring @@ -200,14 +200,11 @@ funext h.coe_re_apply_self /-- A matrix is hermitian iff the corresponding linear map is self adjoint. -/ lemma is_hermitian_iff_is_symmetric [fintype n] [decidable_eq n] {A : matrix n n α} : - is_hermitian A ↔ linear_map.is_symmetric - ((pi_Lp.linear_equiv 2 α (λ _ : n, α)).symm.conj A.to_lin' : module.End α (pi_Lp 2 _)) := + is_hermitian A ↔ A.to_euclidean_lin.is_symmetric := begin rw [linear_map.is_symmetric, (pi_Lp.equiv 2 (λ _ : n, α)).symm.surjective.forall₂], - simp only [linear_equiv.conj_apply, linear_map.comp_apply, linear_equiv.coe_coe, - pi_Lp.linear_equiv_apply, pi_Lp.linear_equiv_symm_apply, linear_equiv.symm_symm], - simp_rw [euclidean_space.inner_eq_star_dot_product, equiv.apply_symm_apply, to_lin'_apply, - star_mul_vec, dot_product_mul_vec], + simp only [to_euclidean_lin_pi_Lp_equiv_symm, euclidean_space.inner_pi_Lp_equiv_symm, + to_lin'_apply, star_mul_vec, dot_product_mul_vec], split, { rintro (h : Aᴴ = A) x y, rw h }, diff --git a/src/linear_algebra/matrix/ldl.lean b/src/linear_algebra/matrix/ldl.lean index 1f0f2f720dd3b..d21b4263db629 100644 --- a/src/linear_algebra/matrix/ldl.lean +++ b/src/linear_algebra/matrix/ldl.lean @@ -30,8 +30,7 @@ decomposed as `S = LDLᴴ` where `L` is a lower-triangular matrix and `D` is a d variables {𝕜 : Type*} [is_R_or_C 𝕜] variables {n : Type*} [linear_order n] [is_well_order n (<)] [locally_finite_order_bot n] -local notation `⟪`x`, `y`⟫` := -@inner 𝕜 (n → 𝕜) (pi_Lp.inner_product_space (λ _, 𝕜)).to_has_inner x y +local notation `⟪`x`, `y`⟫ₑ` := @inner 𝕜 _ _ ((pi_Lp.equiv 2 _).symm x) ((pi_Lp.equiv _ _).symm y) open matrix open_locale matrix @@ -42,7 +41,7 @@ applying Gram-Schmidt-Orthogonalization w.r.t. the inner product induced by `S basis vectors `pi.basis_fun`. -/ noncomputable def LDL.lower_inv : matrix n n 𝕜 := @gram_schmidt - 𝕜 (n → 𝕜) _ (inner_product_space.of_matrix hS.transpose) n _ _ _ (λ k, pi.basis_fun 𝕜 n k) + 𝕜 (n → 𝕜) _ (inner_product_space.of_matrix hS.transpose) n _ _ _ (pi.basis_fun 𝕜 n) lemma LDL.lower_inv_eq_gram_schmidt_basis : LDL.lower_inv hS = ((pi.basis_fun 𝕜 n).to_matrix @@ -64,7 +63,7 @@ begin end lemma LDL.lower_inv_orthogonal {i j : n} (h₀ : i ≠ j) : - ⟪(LDL.lower_inv hS i), Sᵀ.mul_vec (LDL.lower_inv hS j)⟫ = 0 := + ⟪(LDL.lower_inv hS i), Sᵀ.mul_vec (LDL.lower_inv hS j)⟫ₑ = 0 := show @inner 𝕜 (n → 𝕜) (inner_product_space.of_matrix hS.transpose).to_has_inner (LDL.lower_inv hS i) (LDL.lower_inv hS j) = 0, @@ -72,7 +71,7 @@ by apply gram_schmidt_orthogonal _ _ h₀ /-- The entries of the diagonal matrix `D` of the LDL decomposition. -/ noncomputable def LDL.diag_entries : n → 𝕜 := - λ i, ⟪star (LDL.lower_inv hS i), S.mul_vec (star (LDL.lower_inv hS i))⟫ +λ i, ⟪star (LDL.lower_inv hS i), S.mul_vec (star (LDL.lower_inv hS i))⟫ₑ /-- The diagonal matrix `D` of the LDL decomposition. -/ noncomputable def LDL.diag : matrix n n 𝕜 := matrix.diagonal (LDL.diag_entries hS) @@ -88,12 +87,12 @@ lemma LDL.diag_eq_lower_inv_conj : LDL.diag hS = LDL.lower_inv hS ⬝ S ⬝ (LDL begin ext i j, by_cases hij : i = j, - { simpa only [hij, LDL.diag, diagonal_apply_eq, LDL.diag_entries, matrix.mul_assoc, inner, - pi.star_apply, is_R_or_C.star_def, star_ring_end_self_apply] }, + { simpa only [hij, LDL.diag, diagonal_apply_eq, LDL.diag_entries, matrix.mul_assoc, + euclidean_space.inner_pi_Lp_equiv_symm, star_star] }, { simp only [LDL.diag, hij, diagonal_apply_ne, ne.def, not_false_iff, mul_mul_apply], rw [conj_transpose, transpose_map, transpose_transpose, dot_product_mul_vec, (LDL.lower_inv_orthogonal hS (λ h : j = i, hij h.symm)).symm, - ← inner_conj_symm, mul_vec_transpose, euclidean_space.inner_eq_star_dot_product, + ← inner_conj_symm, mul_vec_transpose, euclidean_space.inner_pi_Lp_equiv_symm, ← is_R_or_C.star_def, ← star_dot_product_star, dot_product_comm, star_star], refl } end diff --git a/src/linear_algebra/matrix/spectrum.lean b/src/linear_algebra/matrix/spectrum.lean index 57f56f3598dcc..632d74e1e7e9b 100644 --- a/src/linear_algebra/matrix/spectrum.lean +++ b/src/linear_algebra/matrix/spectrum.lean @@ -90,18 +90,14 @@ theorem spectral_theorem : begin rw [eigenvector_matrix_inv, pi_Lp.basis_to_matrix_basis_fun_mul], ext i j, - have : linear_map.is_symmetric _ := is_hermitian_iff_is_symmetric.1 hA, + have := is_hermitian_iff_is_symmetric.1 hA, convert this.diagonalization_basis_apply_self_apply finrank_euclidean_space (euclidean_space.single j 1) - ((fintype.equiv_of_card_eq (fintype.card_fin _)).symm i), - { dsimp only [linear_equiv.conj_apply_apply, pi_Lp.linear_equiv_apply, - pi_Lp.linear_equiv_symm_apply, pi_Lp.equiv_single, linear_map.std_basis, - linear_map.coe_single, pi_Lp.equiv_symm_single, linear_equiv.symm_symm, - eigenvector_basis, to_lin'_apply], - simp only [basis.to_matrix, basis.coe_to_orthonormal_basis_repr, basis.equiv_fun_apply], - simp_rw [orthonormal_basis.coe_to_basis_repr_apply, orthonormal_basis.repr_reindex, - linear_equiv.symm_symm, pi_Lp.linear_equiv_apply, pi_Lp.equiv_single, mul_vec_single, - mul_one], + ((fintype.equiv_of_card_eq (fintype.card_fin _)).symm i) using 1, + { dsimp only [euclidean_space.single, to_euclidean_lin_pi_Lp_equiv_symm, to_lin'_apply, + matrix.of_apply, is_hermitian.eigenvector_basis], + simp_rw [mul_vec_single, mul_one, orthonormal_basis.coe_to_basis_repr_apply, + orthonormal_basis.repr_reindex], refl }, { simp only [diagonal_mul, (∘), eigenvalues, eigenvector_basis], rw [basis.to_matrix_apply, From 02ba8949f486ebecf93fe7460f1ed0564b5e442c Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Fri, 17 Mar 2023 14:58:11 +0000 Subject: [PATCH 0650/1291] chore (data/finset/sym): remove unnecessary alias (#18603) See the discussion in this [PR](https://github.com/leanprover-community/mathlib4/pull/2168) --- src/data/finset/sym.lean | 4 ---- 1 file changed, 4 deletions(-) diff --git a/src/data/finset/sym.lean b/src/data/finset/sym.lean index bed5ded31f54b..ec506a8667d3d 100644 --- a/src/data/finset/sym.lean +++ b/src/data/finset/sym.lean @@ -141,10 +141,6 @@ end @[simp] lemma sym_nonempty : (s.sym n).nonempty ↔ n = 0 ∨ s.nonempty := by simp_rw [nonempty_iff_ne_empty, ne.def, sym_eq_empty, not_and_distrib, not_ne_iff] -alias sym_nonempty ↔ _ nonempty.sym - -attribute [protected] nonempty.sym - @[simp] lemma sym_univ [fintype α] (n : ℕ) : (univ : finset α).sym n = univ := eq_univ_iff_forall.2 $ λ s, mem_sym_iff.2 $ λ a _, mem_univ _ From acebd8d49928f6ed8920e502a6c90674e75bd441 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 17 Mar 2023 19:18:02 +0000 Subject: [PATCH 0651/1291] feat(algebra/*/opposite): Missing instances (#18602) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit A few missing instances about `nat.cast`/`int.cast`/`rat.cast` and `mul_opposite`/`add_opposite`. Also add the (weirdly) missing `add_comm_group_with_one → add_comm_monoid_with_one`. Finally, this changes the defeq of `rat.cast` on `mul_opposite` to be simpler. --- src/algebra/field/opposite.lean | 26 ++++++++++++++++++++- src/algebra/group/opposite.lean | 41 +++++++++++++++++++++++++++------ src/algebra/ring/opposite.lean | 3 ++- src/data/int/cast/defs.lean | 5 ++-- src/data/int/cast/lemmas.lean | 9 -------- src/data/nat/cast/basic.lean | 9 -------- src/data/rat/cast.lean | 14 ----------- 7 files changed, 64 insertions(+), 43 deletions(-) diff --git a/src/algebra/field/opposite.lean b/src/algebra/field/opposite.lean index b7c310ef76717..ef1e470581d5a 100644 --- a/src/algebra/field/opposite.lean +++ b/src/algebra/field/opposite.lean @@ -5,6 +5,7 @@ Authors: Kenny Lau -/ import algebra.field.defs import algebra.ring.opposite +import data.int.cast.lemmas /-! # Field structure on the multiplicative/additive opposite @@ -15,11 +16,28 @@ import algebra.ring.opposite variables (α : Type*) +namespace mul_opposite + +@[to_additive] instance [has_rat_cast α] : has_rat_cast αᵐᵒᵖ := ⟨λ n, op n⟩ + +variables {α} + +@[simp, norm_cast, to_additive] +lemma op_rat_cast [has_rat_cast α] (q : ℚ) : op (q : α) = q := rfl + +@[simp, norm_cast, to_additive] +lemma unop_rat_cast [has_rat_cast α] (q : ℚ) : unop (q : αᵐᵒᵖ) = q := rfl + +variables (α) + instance [division_semiring α] : division_semiring αᵐᵒᵖ := { .. mul_opposite.group_with_zero α, .. mul_opposite.semiring α } instance [division_ring α] : division_ring αᵐᵒᵖ := -{ .. mul_opposite.group_with_zero α, .. mul_opposite.ring α } +{ rat_cast := λ q, op q, + rat_cast_mk := λ a b hb h, by { rw [rat.cast_def, op_div, op_nat_cast, op_int_cast], + exact int.commute_cast _ _ }, + ..mul_opposite.division_semiring α, ..mul_opposite.ring α } instance [semifield α] : semifield αᵐᵒᵖ := { .. mul_opposite.division_semiring α, .. mul_opposite.comm_semiring α } @@ -27,6 +45,10 @@ instance [semifield α] : semifield αᵐᵒᵖ := instance [field α] : field αᵐᵒᵖ := { .. mul_opposite.division_ring α, .. mul_opposite.comm_ring α } +end mul_opposite + +namespace add_opposite + instance [division_semiring α] : division_semiring αᵃᵒᵖ := { ..add_opposite.group_with_zero α, ..add_opposite.semiring α } @@ -38,3 +60,5 @@ instance [semifield α] : semifield αᵃᵒᵖ := instance [field α] : field αᵃᵒᵖ := { ..add_opposite.division_ring α, ..add_opposite.comm_ring α } + +end add_opposite diff --git a/src/algebra/group/opposite.lean b/src/algebra/group/opposite.lean index 72b86f1524c91..f5c4d12e8848f 100644 --- a/src/algebra/group/opposite.lean +++ b/src/algebra/group/opposite.lean @@ -24,6 +24,9 @@ namespace mul_opposite ### Additive structures on `αᵐᵒᵖ` -/ +@[to_additive] instance [has_nat_cast α] : has_nat_cast αᵐᵒᵖ := ⟨λ n, op n⟩ +@[to_additive] instance [has_int_cast α] : has_int_cast αᵐᵒᵖ := ⟨λ n, op n⟩ + instance [add_semigroup α] : add_semigroup (αᵐᵒᵖ) := unop_injective.add_semigroup _ (λ x y, rfl) @@ -42,14 +45,16 @@ unop_injective.add_zero_class _ rfl (λ x y, rfl) instance [add_monoid α] : add_monoid αᵐᵒᵖ := unop_injective.add_monoid _ rfl (λ _ _, rfl) (λ _ _, rfl) +instance [add_comm_monoid α] : add_comm_monoid αᵐᵒᵖ := +unop_injective.add_comm_monoid _ rfl (λ _ _, rfl) (λ _ _, rfl) + instance [add_monoid_with_one α] : add_monoid_with_one αᵐᵒᵖ := -{ nat_cast := λ n, op n, - nat_cast_zero := show op ((0 : ℕ) : α) = 0, by simp, +{ nat_cast_zero := show op ((0 : ℕ) : α) = 0, by rw [nat.cast_zero, op_zero], nat_cast_succ := show ∀ n, op ((n + 1 : ℕ) : α) = op (n : ℕ) + 1, by simp, - .. mul_opposite.add_monoid α, .. mul_opposite.has_one α } + .. mul_opposite.add_monoid α, .. mul_opposite.has_one α, ..mul_opposite.has_nat_cast _ } -instance [add_comm_monoid α] : add_comm_monoid αᵐᵒᵖ := -unop_injective.add_comm_monoid _ rfl (λ _ _, rfl) (λ _ _, rfl) +instance [add_comm_monoid_with_one α] : add_comm_monoid_with_one αᵐᵒᵖ := +{ .. mul_opposite.add_monoid_with_one α, ..mul_opposite.add_comm_monoid α } instance [sub_neg_monoid α] : sub_neg_monoid αᵐᵒᵖ := unop_injective.sub_neg_monoid _ rfl (λ _ _, rfl) (λ _, rfl) (λ _ _, rfl) (λ _ _, rfl) (λ _ _, rfl) @@ -57,6 +62,9 @@ unop_injective.sub_neg_monoid _ rfl (λ _ _, rfl) (λ _, rfl) (λ _ _, rfl) (λ instance [add_group α] : add_group αᵐᵒᵖ := unop_injective.add_group _ rfl (λ _ _, rfl) (λ _, rfl) (λ _ _, rfl) (λ _ _, rfl) (λ _ _, rfl) +instance [add_comm_group α] : add_comm_group αᵐᵒᵖ := +unop_injective.add_comm_group _ rfl (λ _ _, rfl) (λ _, rfl) (λ _ _, rfl) (λ _ _, rfl) (λ _ _, rfl) + instance [add_group_with_one α] : add_group_with_one αᵐᵒᵖ := { int_cast := λ n, op n, int_cast_of_nat := λ n, show op ((n : ℤ) : α) = op n, by rw int.cast_coe_nat, @@ -64,8 +72,8 @@ instance [add_group_with_one α] : add_group_with_one αᵐᵒᵖ := by erw [unop_op, int.cast_neg_succ_of_nat]; refl, .. mul_opposite.add_monoid_with_one α, .. mul_opposite.add_group α } -instance [add_comm_group α] : add_comm_group αᵐᵒᵖ := -unop_injective.add_comm_group _ rfl (λ _ _, rfl) (λ _, rfl) (λ _ _, rfl) (λ _ _, rfl) (λ _ _, rfl) +instance [add_comm_group_with_one α] : add_comm_group_with_one αᵐᵒᵖ := +{ .. mul_opposite.add_group_with_one α, ..mul_opposite.add_comm_group α } /-! ### Multiplicative structures on `αᵐᵒᵖ` @@ -142,6 +150,15 @@ We also generate additive structures on `αᵃᵒᵖ` using `to_additive` variable {α} +@[simp, norm_cast, to_additive] lemma op_nat_cast [has_nat_cast α] (n : ℕ) : op (n : α) = n := rfl +@[simp, norm_cast, to_additive] lemma op_int_cast [has_int_cast α] (n : ℤ) : op (n : α) = n := rfl + +@[simp, norm_cast, to_additive] +lemma unop_nat_cast [has_nat_cast α] (n : ℕ) : unop (n : αᵐᵒᵖ) = n := rfl + +@[simp, norm_cast, to_additive] +lemma unop_int_cast [has_int_cast α] (n : ℤ) : unop (n : αᵐᵒᵖ) = n := rfl + @[simp, to_additive] lemma unop_div [div_inv_monoid α] (x y : αᵐᵒᵖ) : unop (x / y) = (unop y)⁻¹ * unop x := rfl @@ -232,6 +249,16 @@ unop_injective.group _ rfl (λ _ _, rfl) (λ _, rfl) (λ _ _, rfl) (λ _ _, rfl) instance [comm_group α] : comm_group αᵃᵒᵖ := unop_injective.comm_group _ rfl (λ _ _, rfl) (λ _, rfl) (λ _ _, rfl) (λ _ _, rfl) (λ _ _, rfl) +-- NOTE: `add_monoid_with_one α → add_monoid_with_one αᵃᵒᵖ` does not hold + +instance [add_comm_monoid_with_one α] : add_comm_monoid_with_one αᵃᵒᵖ := +{ nat_cast_zero := show op ((0 : ℕ) : α) = 0, by rw [nat.cast_zero, op_zero], + nat_cast_succ := show ∀ n, op ((n + 1 : ℕ) : α) = op (n : ℕ) + 1, by simp [add_comm], + ..add_opposite.add_comm_monoid α, ..add_opposite.has_one, ..add_opposite.has_nat_cast _ } + +instance [add_comm_group_with_one α] : add_comm_group_with_one αᵃᵒᵖ := +{ ..add_opposite.add_comm_monoid_with_one _, ..add_opposite.add_comm_group α } + variable {α} /-- The function `add_opposite.op` is a multiplicative equivalence. -/ diff --git a/src/algebra/ring/opposite.lean b/src/algebra/ring/opposite.lean index e9d57dd7a7955..601e9da6b3baf 100644 --- a/src/algebra/ring/opposite.lean +++ b/src/algebra/ring/opposite.lean @@ -123,7 +123,8 @@ instance [non_unital_semiring α] : non_unital_semiring αᵃᵒᵖ := { .. add_opposite.semigroup_with_zero α, .. add_opposite.non_unital_non_assoc_semiring α } instance [non_assoc_semiring α] : non_assoc_semiring αᵃᵒᵖ := -{ .. add_opposite.mul_zero_one_class α, .. add_opposite.non_unital_non_assoc_semiring α } +{ ..add_opposite.mul_zero_one_class α, ..add_opposite.non_unital_non_assoc_semiring α, + ..add_opposite.add_comm_monoid_with_one _ } instance [semiring α] : semiring αᵃᵒᵖ := { .. add_opposite.non_unital_semiring α, .. add_opposite.non_assoc_semiring α, diff --git a/src/data/int/cast/defs.lean b/src/data/int/cast/defs.lean index bcb87ad797ead..ec6e9b20d5463 100644 --- a/src/data/int/cast/defs.lean +++ b/src/data/int/cast/defs.lean @@ -52,8 +52,9 @@ class add_group_with_one (R : Type u) (int_cast_neg_succ_of_nat : ∀ n : ℕ, int_cast (-(n+1 : ℕ)) = -((n+1 : ℕ) : R) . control_laws_tac) /-- An `add_comm_group_with_one` is an `add_group_with_one` satisfying `a + b = b + a`. -/ -@[protect_proj, ancestor add_comm_group add_group_with_one] -class add_comm_group_with_one (R : Type u) extends add_comm_group R, add_group_with_one R +@[protect_proj, ancestor add_comm_group add_group_with_one add_comm_monoid_with_one] +class add_comm_group_with_one (R : Type u) + extends add_comm_group R, add_group_with_one R, add_comm_monoid_with_one R /-- Canonical homomorphism from the integers to any ring(-like) structure `R` -/ protected def int.cast {R : Type u} [has_int_cast R] (i : ℤ) : R := has_int_cast.int_cast i diff --git a/src/data/int/cast/lemmas.lean b/src/data/int/cast/lemmas.lean index 75356b0146c7e..88b021a9cab0c 100644 --- a/src/data/int/cast/lemmas.lean +++ b/src/data/int/cast/lemmas.lean @@ -290,15 +290,6 @@ by refine_struct { .. }; tactic.pi_instance_derive_field end pi -namespace mul_opposite -variables [add_group_with_one α] - -@[simp, norm_cast] lemma op_int_cast (z : ℤ) : op (z : α) = z := rfl - -@[simp, norm_cast] lemma unop_int_cast (n : ℤ) : unop (n : αᵐᵒᵖ) = n := rfl - -end mul_opposite - /-! ### Order dual -/ open order_dual diff --git a/src/data/nat/cast/basic.lean b/src/data/nat/cast/basic.lean index adee13377c6d0..1f06d0878e807 100644 --- a/src/data/nat/cast/basic.lean +++ b/src/data/nat/cast/basic.lean @@ -223,15 +223,6 @@ rfl instance nat.unique_ring_hom {R : Type*} [non_assoc_semiring R] : unique (ℕ →+* R) := { default := nat.cast_ring_hom R, uniq := ring_hom.eq_nat_cast' } -namespace mul_opposite -variables [add_monoid_with_one α] - -@[simp, norm_cast] lemma op_nat_cast (n : ℕ) : op (n : α) = n := rfl - -@[simp, norm_cast] lemma unop_nat_cast (n : ℕ) : unop (n : αᵐᵒᵖ) = n := rfl - -end mul_opposite - namespace pi variables {π : α → Type*} [Π a, has_nat_cast (π a)] diff --git a/src/data/rat/cast.lean b/src/data/rat/cast.lean index 50bd3848cea21..d4040ea4f3b16 100644 --- a/src/data/rat/cast.lean +++ b/src/data/rat/cast.lean @@ -310,20 +310,6 @@ monoid_with_zero_hom.ext_rat' $ ring_hom.congr_fun $ instance rat.subsingleton_ring_hom {R : Type*} [semiring R] : subsingleton (ℚ →+* R) := ⟨ring_hom.ext_rat⟩ -namespace mul_opposite - -variables [division_ring α] - -@[simp, norm_cast] lemma op_rat_cast (r : ℚ) : op (r : α) = (↑r : αᵐᵒᵖ) := -by rw [cast_def, div_eq_mul_inv, op_mul, op_inv, op_nat_cast, op_int_cast, - (commute.cast_int_right _ r.num).eq, cast_def, div_eq_mul_inv] - -@[simp, norm_cast] lemma unop_rat_cast (r : ℚ) : unop (r : αᵐᵒᵖ) = r := -by rw [cast_def, div_eq_mul_inv, unop_mul, unop_inv, unop_nat_cast, unop_int_cast, - (commute.cast_int_right _ r.num).eq, cast_def, div_eq_mul_inv] - -end mul_opposite - section smul namespace rat From 20715f4ac6819ef2453d9e5106ecd086a5dc2a5e Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 18 Mar 2023 01:10:33 +0000 Subject: [PATCH 0652/1291] feat(data/set/sups): Set family operations (#18172) Followup to #17947. Add a similar `set` API (but do not define `set.disj_sups` because I don't need it), correct a few lemma names and connect to upper/lower sets. --- src/data/finset/n_ary.lean | 17 ++- src/data/finset/sups.lean | 110 +++++++++++-------- src/data/set/n_ary.lean | 18 ++- src/data/set/sups.lean | 220 +++++++++++++++++++++++++++++++++++++ 4 files changed, 315 insertions(+), 50 deletions(-) create mode 100644 src/data/set/sups.lean diff --git a/src/data/finset/n_ary.lean b/src/data/finset/n_ary.lean index 34cef362469ca..80aa9c093ba97 100644 --- a/src/data/finset/n_ary.lean +++ b/src/data/finset/n_ary.lean @@ -27,7 +27,7 @@ open function set namespace finset -variables {α α' β β' γ γ' δ δ' ε ε' : Type*} +variables {α α' β β' γ γ' δ δ' ε ε' ζ ζ' ν : Type*} [decidable_eq α'] [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 γ} @@ -71,11 +71,11 @@ lemma image₂_subset_left (ht : t ⊆ t') : image₂ f s t ⊆ image₂ f s t' 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₂_left (hb : b ∈ t) : s.image (λ a, f a b) ⊆ image₂ f s t := +image_subset_iff.2 $ λ 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 image_subset_image₂_right (ha : a ∈ s) : t.image (f a) ⊆ image₂ f s t := +image_subset_iff.2 $ λ 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] @@ -252,6 +252,13 @@ lemma image₂_right_comm {γ : Type*} {u : finset γ} {f : δ → γ → ε} {g 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₂_image₂_comm {γ δ : Type*} {u : finset γ} {v : finset δ} [decidable_eq ζ] + [decidable_eq ζ'] [decidable_eq ν] {f : ε → ζ → ν} {g : α → β → ε} {h : γ → δ → ζ} + {f' : ε' → ζ' → ν} {g' : α → γ → ε'} {h' : β → δ → ζ'} + (h_comm : ∀ a b c d, f (g a b) (h c d) = f' (g' a c) (h' b d)) : + image₂ f (image₂ g s t) (image₂ h u v) = image₂ f' (image₂ g' s u) (image₂ h' t v) := +coe_injective $ by { push_cast, exact image2_image2_image2_comm h_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₂) := diff --git a/src/data/finset/sups.lean b/src/data/finset/sups.lean index 97f93b29855e2..6928f2d00bc6e 100644 --- a/src/data/finset/sups.lean +++ b/src/data/finset/sups.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ import data.finset.n_ary +import data.set.sups /-! # Set family operations @@ -12,16 +13,16 @@ This file defines a few binary operations on `finset α` for use in set family c ## Main declarations -* `finset.sups s t`: Finset of elements of the form `a ⊔ b` where `a ∈ s`, `b ∈ t`. -* `finset.infs s t`: Finset of elements of the form `a ⊓ b` where `a ∈ s`, `b ∈ t`. +* `s ⊻ t`: Finset of elements of the form `a ⊔ b` where `a ∈ s`, `b ∈ t`. +* `s ⊼ t`: Finset of elements of the form `a ⊓ b` where `a ∈ s`, `b ∈ t`. * `finset.disj_sups s t`: Finset of elements of the form `a ⊔ b` where `a ∈ s`, `b ∈ t` and `a` and `b` are disjoint. ## Notation We define the following notation in locale `finset_family`: -* `s ⊻ t` for `finset.sups s t` -* `s ⊼ t` for `finset.infs s t` +* `s ⊻ t` +* `s ⊼ t` * `s ○ t` for `finset.disj_sups s t` ## References @@ -30,25 +31,26 @@ We define the following notation in locale `finset_family`: -/ open function +open_locale set_family variables {α : Type*} [decidable_eq α] namespace finset section sups -variables [semilattice_sup α] (s s₁ s₂ t t₁ t₂ u : finset α) +variables [semilattice_sup α] (s s₁ s₂ t t₁ t₂ u v : finset α) -/-- The finset of elements of the form `a ⊔ b` where `a ∈ s`, `b ∈ t`. -/ -def sups (s t : finset α) : finset α := image₂ (⊔) s t +/-- `s ⊻ t` is the finset of elements of the form `a ⊔ b` where `a ∈ s`, `b ∈ t`. -/ +protected def has_sups : has_sups (finset α) := ⟨image₂ (⊔)⟩ -localized "infix (name := finset.sups) ` ⊻ `:74 := finset.sups" in finset_family +localized "attribute [instance] finset.has_sups" in finset_family variables {s t} {a b c : α} -@[simp] lemma mem_sups : c ∈ s ⊻ t ↔ ∃ (a ∈ s) (b ∈ t), a ⊔ b = c := by simp [sups] +@[simp] lemma mem_sups : c ∈ s ⊻ t ↔ ∃ (a ∈ s) (b ∈ t), a ⊔ b = c := by simp [(⊻)] variables (s t) -@[simp, norm_cast] lemma coe_sups : (s ⊻ t : set α) = set.image2 (⊔) s t := coe_image₂ _ _ _ +@[simp, norm_cast] lemma coe_sups : (↑(s ⊻ t) : set α) = s ⊻ t := coe_image₂ _ _ _ lemma card_sups_le : (s ⊻ t).card ≤ s.card * t.card := card_image₂_le _ _ _ @@ -63,29 +65,28 @@ lemma sups_subset : s₁ ⊆ s₂ → t₁ ⊆ t₂ → s₁ ⊻ t₁ ⊆ s₂ lemma sups_subset_left : t₁ ⊆ t₂ → s ⊻ t₁ ⊆ s ⊻ t₂ := image₂_subset_left lemma sups_subset_right : s₁ ⊆ s₂ → s₁ ⊻ t ⊆ s₂ ⊻ t := image₂_subset_right -lemma image_subset_sups_left : b ∈ t → (λ a, a ⊔ b) '' s ⊆ s ⊻ t := image_subset_image₂_left -lemma image_subset_sups_right : a ∈ s → (⊔) a '' t ⊆ s ⊻ t := image_subset_image₂_right +lemma image_subset_sups_left : b ∈ t → s.image (λ a, a ⊔ b) ⊆ s ⊻ t := image_subset_image₂_left +lemma image_subset_sups_right : a ∈ s → t.image ((⊔) a) ⊆ s ⊻ t := image_subset_image₂_right lemma forall_sups_iff {p : α → Prop} : (∀ c ∈ s ⊻ t, p c) ↔ ∀ (a ∈ s) (b ∈ t), p (a ⊔ b) := forall_image₂_iff @[simp] lemma sups_subset_iff : s ⊻ t ⊆ u ↔ ∀ (a ∈ s) (b ∈ t), a ⊔ b ∈ u := image₂_subset_iff -@[simp] lemma sups_nonempty_iff : (s ⊻ t).nonempty ↔ s.nonempty ∧ t.nonempty := image₂_nonempty_iff +@[simp] lemma sups_nonempty : (s ⊻ t).nonempty ↔ s.nonempty ∧ t.nonempty := image₂_nonempty_iff -lemma nonempty.sups : s.nonempty → t.nonempty → (s ⊻ t).nonempty := nonempty.image₂ +protected lemma nonempty.sups : s.nonempty → t.nonempty → (s ⊻ t).nonempty := nonempty.image₂ lemma nonempty.of_sups_left : (s ⊻ t).nonempty → s.nonempty := nonempty.of_image₂_left lemma nonempty.of_sups_right : (s ⊻ t).nonempty → t.nonempty := nonempty.of_image₂_right -@[simp] lemma sups_empty_left : ∅ ⊻ t = ∅ := image₂_empty_left -@[simp] lemma sups_empty_right : s ⊻ ∅ = ∅ := image₂_empty_right -@[simp] lemma sups_eq_empty_iff : s ⊻ t = ∅ ↔ s = ∅ ∨ t = ∅ := image₂_eq_empty_iff +@[simp] lemma empty_sups : ∅ ⊻ t = ∅ := image₂_empty_left +@[simp] lemma sups_empty : s ⊻ ∅ = ∅ := image₂_empty_right +@[simp] lemma sups_eq_empty : s ⊻ t = ∅ ↔ s = ∅ ∨ t = ∅ := image₂_eq_empty_iff -@[simp] lemma sups_singleton_left : {a} ⊻ t = t.image (λ b, a ⊔ b) := image₂_singleton_left -@[simp] lemma sups_singleton_right : s ⊻ {b} = s.image (λ a, a ⊔ b) := image₂_singleton_right -lemma sups_singleton_left' : {a} ⊻ t = t.image ((⊔) a) := image₂_singleton_left' +@[simp] lemma singleton_sups : {a} ⊻ t = t.image (λ b, a ⊔ b) := image₂_singleton_left +@[simp] lemma sups_singleton : s ⊻ {b} = s.image (λ a, a ⊔ b) := image₂_singleton_right -lemma sups_singleton : ({a} ⊻ {b} : finset α) = {a ⊔ b} := image₂_singleton +lemma singleton_sups_singleton : ({a} ⊻ {b} : finset α) = {a ⊔ b} := image₂_singleton lemma sups_union_left : (s₁ ∪ s₂) ⊻ t = s₁ ⊻ t ∪ s₂ ⊻ t := image₂_union_left lemma sups_union_right : s ⊻ (t₁ ∪ t₂) = s ⊻ t₁ ∪ s ⊻ t₂ := image₂_union_right @@ -94,10 +95,10 @@ lemma sups_inter_subset_left : (s₁ ∩ s₂) ⊻ t ⊆ s₁ ⊻ t ∩ s₂ ⊻ lemma sups_inter_subset_right : s ⊻ (t₁ ∩ t₂) ⊆ s ⊻ t₁ ∩ s ⊻ t₂ := image₂_inter_subset_right lemma subset_sups {s t : set α} : - ↑u ⊆ set.image2 (⊔) s t → ∃ s' t' : finset α, ↑s' ⊆ s ∧ ↑t' ⊆ t ∧ u ⊆ s' ⊻ t' := + ↑u ⊆ s ⊻ t → ∃ s' t' : finset α, ↑s' ⊆ s ∧ ↑t' ⊆ t ∧ u ⊆ s' ⊻ t' := subset_image₂ -variables (s t u) +variables (s t u v) lemma bUnion_image_sup_left : s.bUnion (λ a, t.image $ (⊔) a) = s ⊻ t := bUnion_image_left lemma bUnion_image_sup_right : t.bUnion (λ b, s.image $ λ a, a ⊔ b) = s ⊻ t := bUnion_image_right @@ -109,24 +110,26 @@ lemma sups_assoc : (s ⊻ t) ⊻ u = s ⊻ (t ⊻ u) := image₂_assoc $ λ _ _ lemma sups_comm : s ⊻ t = t ⊻ s := image₂_comm $ λ _ _, sup_comm lemma sups_left_comm : s ⊻ (t ⊻ u) = t ⊻ (s ⊻ u) := image₂_left_comm sup_left_comm lemma sups_right_comm : (s ⊻ t) ⊻ u = (s ⊻ u) ⊻ t := image₂_right_comm sup_right_comm +lemma sups_sups_sups_comm : (s ⊻ t) ⊻ (u ⊻ v) = (s ⊻ u) ⊻ (t ⊻ v) := +image₂_image₂_image₂_comm sup_sup_sup_comm end sups section infs -variables [semilattice_inf α] (s s₁ s₂ t t₁ t₂ u : finset α) +variables [semilattice_inf α] (s s₁ s₂ t t₁ t₂ u v : finset α) -/-- The finset of elements of the form `a ⊓ b` where `a ∈ s`, `b ∈ t`. -/ -def infs (s t : finset α) : finset α := image₂ (⊓) s t +/-- `s ⊼ t` is the finset of elements of the form `a ⊓ b` where `a ∈ s`, `b ∈ t`. -/ +protected def has_infs : has_infs (finset α) := ⟨image₂ (⊓)⟩ -localized "infix (name := finset.infs) ` ⊼ `:74 := finset.infs" in finset_family +localized "attribute [instance] finset.has_infs" in finset_family variables {s t} {a b c : α} -@[simp] lemma mem_infs : c ∈ s ⊼ t ↔ ∃ (a ∈ s) (b ∈ t), a ⊓ b = c := by simp [infs] +@[simp] lemma mem_infs : c ∈ s ⊼ t ↔ ∃ (a ∈ s) (b ∈ t), a ⊓ b = c := by simp [(⊼)] variables (s t) -@[simp, norm_cast] lemma coe_infs : (s ⊼ t : set α) = set.image2 (⊓) s t := coe_image₂ _ _ _ +@[simp, norm_cast] lemma coe_infs : (↑(s ⊼ t) : set α) = s ⊼ t := coe_image₂ _ _ _ lemma card_infs_le : (s ⊼ t).card ≤ s.card * t.card := card_image₂_le _ _ _ @@ -141,29 +144,28 @@ lemma infs_subset : s₁ ⊆ s₂ → t₁ ⊆ t₂ → s₁ ⊼ t₁ ⊆ s₂ lemma infs_subset_left : t₁ ⊆ t₂ → s ⊼ t₁ ⊆ s ⊼ t₂ := image₂_subset_left lemma infs_subset_right : s₁ ⊆ s₂ → s₁ ⊼ t ⊆ s₂ ⊼ t := image₂_subset_right -lemma image_subset_infs_left : b ∈ t → (λ a, a ⊓ b) '' s ⊆ s ⊼ t := image_subset_image₂_left -lemma image_subset_infs_right : a ∈ s → (⊓) a '' t ⊆ s ⊼ t := image_subset_image₂_right +lemma image_subset_infs_left : b ∈ t → s.image (λ a, a ⊓ b) ⊆ s ⊼ t := image_subset_image₂_left +lemma image_subset_infs_right : a ∈ s → t.image ((⊓) a) ⊆ s ⊼ t := image_subset_image₂_right lemma forall_infs_iff {p : α → Prop} : (∀ c ∈ s ⊼ t, p c) ↔ ∀ (a ∈ s) (b ∈ t), p (a ⊓ b) := forall_image₂_iff @[simp] lemma infs_subset_iff : s ⊼ t ⊆ u ↔ ∀ (a ∈ s) (b ∈ t), a ⊓ b ∈ u := image₂_subset_iff -@[simp] lemma infs_nonempty_iff : (s ⊼ t).nonempty ↔ s.nonempty ∧ t.nonempty := image₂_nonempty_iff +@[simp] lemma infs_nonempty : (s ⊼ t).nonempty ↔ s.nonempty ∧ t.nonempty := image₂_nonempty_iff -lemma nonempty.infs : s.nonempty → t.nonempty → (s ⊼ t).nonempty := nonempty.image₂ +protected lemma nonempty.infs : s.nonempty → t.nonempty → (s ⊼ t).nonempty := nonempty.image₂ lemma nonempty.of_infs_left : (s ⊼ t).nonempty → s.nonempty := nonempty.of_image₂_left lemma nonempty.of_infs_right : (s ⊼ t).nonempty → t.nonempty := nonempty.of_image₂_right -@[simp] lemma infs_empty_left : ∅ ⊼ t = ∅ := image₂_empty_left -@[simp] lemma infs_empty_right : s ⊼ ∅ = ∅ := image₂_empty_right -@[simp] lemma infs_eq_empty_iff : s ⊼ t = ∅ ↔ s = ∅ ∨ t = ∅ := image₂_eq_empty_iff +@[simp] lemma empty_infs : ∅ ⊼ t = ∅ := image₂_empty_left +@[simp] lemma infs_empty : s ⊼ ∅ = ∅ := image₂_empty_right +@[simp] lemma infs_eq_empty : s ⊼ t = ∅ ↔ s = ∅ ∨ t = ∅ := image₂_eq_empty_iff -@[simp] lemma infs_singleton_left : {a} ⊼ t = t.image (λ b, a ⊓ b) := image₂_singleton_left -@[simp] lemma infs_singleton_right : s ⊼ {b} = s.image (λ a, a ⊓ b) := image₂_singleton_right -lemma infs_singleton_left' : {a} ⊼ t = t.image ((⊓) a) := image₂_singleton_left' +@[simp] lemma singleton_infs : {a} ⊼ t = t.image (λ b, a ⊓ b) := image₂_singleton_left +@[simp] lemma infs_singleton : s ⊼ {b} = s.image (λ a, a ⊓ b) := image₂_singleton_right -lemma infs_singleton : ({a} ⊼ {b} : finset α) = {a ⊓ b} := image₂_singleton +lemma singleton_infs_singleton : ({a} ⊼ {b} : finset α) = {a ⊓ b} := image₂_singleton lemma infs_union_left : (s₁ ∪ s₂) ⊼ t = s₁ ⊼ t ∪ s₂ ⊼ t := image₂_union_left lemma infs_union_right : s ⊼ (t₁ ∪ t₂) = s ⊼ t₁ ∪ s ⊼ t₂ := image₂_union_right @@ -172,10 +174,10 @@ lemma infs_inter_subset_left : (s₁ ∩ s₂) ⊼ t ⊆ s₁ ⊼ t ∩ s₂ ⊼ lemma infs_inter_subset_right : s ⊼ (t₁ ∩ t₂) ⊆ s ⊼ t₁ ∩ s ⊼ t₂ := image₂_inter_subset_right lemma subset_infs {s t : set α} : - ↑u ⊆ set.image2 (⊓) s t → ∃ s' t' : finset α, ↑s' ⊆ s ∧ ↑t' ⊆ t ∧ u ⊆ s' ⊼ t' := + ↑u ⊆ s ⊼ t → ∃ s' t' : finset α, ↑s' ⊆ s ∧ ↑t' ⊆ t ∧ u ⊆ s' ⊼ t' := subset_image₂ -variables (s t u) +variables (s t u v) lemma bUnion_image_inf_left : s.bUnion (λ a, t.image $ (⊓) a) = s ⊼ t := bUnion_image_left lemma bUnion_image_inf_right : t.bUnion (λ b, s.image $ λ a, a ⊓ b) = s ⊼ t := bUnion_image_right @@ -187,11 +189,30 @@ lemma infs_assoc : (s ⊼ t) ⊼ u = s ⊼ (t ⊼ u) := image₂_assoc $ λ _ _ lemma infs_comm : s ⊼ t = t ⊼ s := image₂_comm $ λ _ _, inf_comm lemma infs_left_comm : s ⊼ (t ⊼ u) = t ⊼ (s ⊼ u) := image₂_left_comm inf_left_comm lemma infs_right_comm : (s ⊼ t) ⊼ u = (s ⊼ u) ⊼ t := image₂_right_comm inf_right_comm +lemma infs_infs_infs_comm : (s ⊼ t) ⊼ (u ⊼ v) = (s ⊼ u) ⊼ (t ⊼ v) := +image₂_image₂_image₂_comm inf_inf_inf_comm end infs open_locale finset_family +section distrib_lattice +variables [distrib_lattice α] (s t u : finset α) + +lemma sups_infs_subset_left : s ⊻ (t ⊼ u) ⊆ (s ⊻ t) ⊼ (s ⊻ u) := +image₂_distrib_subset_left $ λ _ _ _, sup_inf_left + +lemma sups_infs_subset_right : (t ⊼ u) ⊻ s ⊆ (t ⊻ s) ⊼ (u ⊻ s) := +image₂_distrib_subset_right $ λ _ _ _, sup_inf_right + +lemma infs_sups_subset_left : s ⊼ (t ⊻ u) ⊆ (s ⊼ t) ⊻ (s ⊼ u) := +image₂_distrib_subset_left $ λ _ _ _, inf_sup_left + +lemma infs_sups_subset_right : (t ⊻ u) ⊼ s ⊆ (t ⊼ s) ⊻ (u ⊼ s) := +image₂_distrib_subset_right $ λ _ _ _, inf_sup_right + +end distrib_lattice + section disj_sups variables [semilattice_sup α] [order_bot α] [@decidable_rel α disjoint] (s s₁ s₂ t t₁ t₂ u : finset α) @@ -271,7 +292,7 @@ end disj_sups open_locale finset_family section distrib_lattice -variables [distrib_lattice α] [order_bot α] [@decidable_rel α disjoint] (s t u : finset α) +variables [distrib_lattice α] [order_bot α] [@decidable_rel α disjoint] (s t u v : finset α) lemma disj_sups_assoc : ∀ s t u : finset α, (s ○ t) ○ u = s ○ (t ○ u) := begin @@ -288,5 +309,8 @@ by simp_rw [←disj_sups_assoc, disj_sups_comm s] lemma disj_sups_right_comm : (s ○ t) ○ u = (s ○ u) ○ t := by simp_rw [disj_sups_assoc, disj_sups_comm] +lemma disj_sups_disj_sups_disj_sups_comm : (s ○ t) ○ (u ○ v) = (s ○ u) ○ (t ○ v) := +by simp_rw [←disj_sups_assoc, disj_sups_right_comm] + end distrib_lattice end finset diff --git a/src/data/set/n_ary.lean b/src/data/set/n_ary.lean index 47d952d27fd72..6b0413ebb416a 100644 --- a/src/data/set/n_ary.lean +++ b/src/data/set/n_ary.lean @@ -26,8 +26,9 @@ and `set.image2` already fulfills this task. open function namespace set -variables {α α' β β' γ γ' δ δ' ε ε' : Type*} {f f' : α → β → γ} {g g' : α → β → γ → δ} -variables {s s' : set α} {t t' : set β} {u u' : set γ} {a a' : α} {b b' : β} {c c' : γ} {d d' : δ} +variables {α α' β β' γ γ' δ δ' ε ε' ζ ζ' ν : Type*} {f f' : α → β → γ} {g g' : α → β → γ → δ} +variables {s s' : set α} {t t' : set β} {u u' : set γ} {v : set δ} {a a' : α} {b b' : β} {c c' : γ} + {d d' : δ} /-- The image of a binary function `f : α → β → γ` as a function `set α → set β → set γ`. @@ -235,6 +236,19 @@ lemma image2_right_comm {f : δ → γ → ε} {g : α → β → δ} {f' : α image2 f (image2 g s t) u = image2 g' (image2 f' s u) t := by { rw [image2_swap g, image2_swap g'], exact image2_assoc (λ _ _ _, h_right_comm _ _ _) } +lemma image2_image2_image2_comm {f : ε → ζ → ν} {g : α → β → ε} {h : γ → δ → ζ} {f' : ε' → ζ' → ν} + {g' : α → γ → ε'} {h' : β → δ → ζ'} + (h_comm : ∀ a b c d, f (g a b) (h c d) = f' (g' a c) (h' b d)) : + image2 f (image2 g s t) (image2 h u v) = image2 f' (image2 g' s u) (image2 h' t v) := +begin + ext, + split, + { rintro ⟨_, _, ⟨a, b, ha, hb, rfl⟩, ⟨c, d, hc, hd, rfl⟩, rfl⟩, + exact ⟨_, _, ⟨a, c, ha, hc, rfl⟩, ⟨b, d, hb, hd, rfl⟩, (h_comm _ _ _ _).symm⟩ }, + { rintro ⟨_, _, ⟨a, c, ha, hc, rfl⟩, ⟨b, d, hb, hd, rfl⟩, rfl⟩, + exact ⟨_, _, ⟨a, b, ha, hb, rfl⟩, ⟨c, d, hc, hd, rfl⟩, h_comm _ _ _ _⟩ } +end + lemma image_image2_distrib {g : γ → δ} {f' : α' → β' → δ} {g₁ : α → α'} {g₂ : β → β'} (h_distrib : ∀ a b, g (f a b) = f' (g₁ a) (g₂ b)) : (image2 f s t).image g = image2 f' (s.image g₁) (t.image g₂) := diff --git a/src/data/set/sups.lean b/src/data/set/sups.lean new file mode 100644 index 0000000000000..d06e854a18b3e --- /dev/null +++ b/src/data/set/sups.lean @@ -0,0 +1,220 @@ +/- +Copyright (c) 2022 Yaël Dillies. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies +-/ +import data.set.n_ary +import order.upper_lower.basic + +/-! +# Set family operations + +This file defines a few binary operations on `set α` for use in set family combinatorics. + +## Main declarations + +* `s ⊻ t`: Set of elements of the form `a ⊔ b` where `a ∈ s`, `b ∈ t`. +* `s ⊼ t`: Set of elements of the form `a ⊓ b` where `a ∈ s`, `b ∈ t`. + +## Notation + +We define the following notation in locale `set_family`: +* `s ⊻ t` +* `s ⊼ t` + +## References + +[B. Bollobás, *Combinatorics*][bollobas1986] +-/ + +open function + +variables {α : Type*} + +/-- Notation typeclass for pointwise supremum `⊻`. -/ +class has_sups (α : Type*) := +(sups : α → α → α) + +/-- Notation typeclass for pointwise infimum `⊼`. -/ +class has_infs (α : Type*) := +(infs : α → α → α) + +-- This notation is meant to have higher precedence than `⊔` and `⊓`, but still within the realm of +-- other binary operations +infix ` ⊻ `:74 := has_sups.sups +infix ` ⊼ `:75 := has_infs.infs + +namespace set +section sups +variables [semilattice_sup α] (s s₁ s₂ t t₁ t₂ u v : set α) + +/-- `s ⊻ t` is the set of elements of the form `a ⊔ b` where `a ∈ s`, `b ∈ t`. -/ +protected def has_sups : has_sups (set α) := ⟨image2 (⊔)⟩ + +localized "attribute [instance] set.has_sups" in set_family + +variables {s s₁ s₂ t t₁ t₂ u} {a b c : α} + +@[simp] lemma mem_sups : c ∈ s ⊻ t ↔ ∃ (a ∈ s) (b ∈ t), a ⊔ b = c := by simp [(⊻)] + +lemma sup_mem_sups : a ∈ s → b ∈ t → a ⊔ b ∈ s ⊻ t := mem_image2_of_mem +lemma sups_subset : s₁ ⊆ s₂ → t₁ ⊆ t₂ → s₁ ⊻ t₁ ⊆ s₂ ⊻ t₂ := image2_subset +lemma sups_subset_left : t₁ ⊆ t₂ → s ⊻ t₁ ⊆ s ⊻ t₂ := image2_subset_left +lemma sups_subset_right : s₁ ⊆ s₂ → s₁ ⊻ t ⊆ s₂ ⊻ t := image2_subset_right + +lemma image_subset_sups_left : b ∈ t → (λ a, a ⊔ b) '' s ⊆ s ⊻ t := image_subset_image2_left +lemma image_subset_sups_right : a ∈ s → (⊔) a '' t ⊆ s ⊻ t := image_subset_image2_right + +lemma forall_sups_iff {p : α → Prop} : (∀ c ∈ s ⊻ t, p c) ↔ ∀ (a ∈ s) (b ∈ t), p (a ⊔ b) := +forall_image2_iff + +@[simp] lemma sups_subset_iff : s ⊻ t ⊆ u ↔ ∀ (a ∈ s) (b ∈ t), a ⊔ b ∈ u := image2_subset_iff + +@[simp] lemma sups_nonempty : (s ⊻ t).nonempty ↔ s.nonempty ∧ t.nonempty := image2_nonempty_iff + +protected lemma nonempty.sups : s.nonempty → t.nonempty → (s ⊻ t).nonempty := nonempty.image2 +lemma nonempty.of_sups_left : (s ⊻ t).nonempty → s.nonempty := nonempty.of_image2_left +lemma nonempty.of_sups_right : (s ⊻ t).nonempty → t.nonempty := nonempty.of_image2_right + +@[simp] lemma empty_sups : ∅ ⊻ t = ∅ := image2_empty_left +@[simp] lemma sups_empty : s ⊻ ∅ = ∅ := image2_empty_right +@[simp] lemma sups_eq_empty : s ⊻ t = ∅ ↔ s = ∅ ∨ t = ∅ := image2_eq_empty_iff + +@[simp] lemma singleton_sups : {a} ⊻ t = t.image (λ b, a ⊔ b) := image2_singleton_left +@[simp] lemma sups_singleton : s ⊻ {b} = s.image (λ a, a ⊔ b) := image2_singleton_right + +lemma singleton_sups_singleton : ({a} ⊻ {b} : set α) = {a ⊔ b} := image2_singleton + +lemma sups_union_left : (s₁ ∪ s₂) ⊻ t = s₁ ⊻ t ∪ s₂ ⊻ t := image2_union_left +lemma sups_union_right : s ⊻ (t₁ ∪ t₂) = s ⊻ t₁ ∪ s ⊻ t₂ := image2_union_right + +lemma sups_inter_subset_left : (s₁ ∩ s₂) ⊻ t ⊆ s₁ ⊻ t ∩ s₂ ⊻ t := image2_inter_subset_left +lemma sups_inter_subset_right : s ⊻ (t₁ ∩ t₂) ⊆ s ⊻ t₁ ∩ s ⊻ t₂ := image2_inter_subset_right + +variables (s t u v) + +lemma Union_image_sup_left : (⋃ a ∈ s, (⊔) a '' t) = s ⊻ t := Union_image_left _ +lemma Union_image_sup_right : (⋃ b ∈ t, (⊔ b) '' s) = s ⊻ t := Union_image_right _ + +@[simp] lemma image_sup_prod (s t : set α) : (s ×ˢ t).image (uncurry (⊔)) = s ⊻ t := +image_uncurry_prod _ _ _ + +lemma sups_assoc : (s ⊻ t) ⊻ u = s ⊻ (t ⊻ u) := image2_assoc $ λ _ _ _, sup_assoc +lemma sups_comm : s ⊻ t = t ⊻ s := image2_comm $ λ _ _, sup_comm +lemma sups_left_comm : s ⊻ (t ⊻ u) = t ⊻ (s ⊻ u) := image2_left_comm sup_left_comm +lemma sups_right_comm : (s ⊻ t) ⊻ u = (s ⊻ u) ⊻ t := image2_right_comm sup_right_comm +lemma sups_sups_sups_comm : (s ⊻ t) ⊻ (u ⊻ v) = (s ⊻ u) ⊻ (t ⊻ v) := +image2_image2_image2_comm sup_sup_sup_comm + +end sups + +section infs +variables [semilattice_inf α] (s s₁ s₂ t t₁ t₂ u v : set α) + +/-- `s ⊼ t` is the set of elements of the form `a ⊓ b` where `a ∈ s`, `b ∈ t`. -/ +protected def has_infs : has_infs (set α) := ⟨image2 (⊓)⟩ + +localized "attribute [instance] set.has_infs" in set_family + +variables {s s₁ s₂ t t₁ t₂ u} {a b c : α} + +@[simp] lemma mem_infs : c ∈ s ⊼ t ↔ ∃ (a ∈ s) (b ∈ t), a ⊓ b = c := by simp [(⊼)] + +lemma inf_mem_infs : a ∈ s → b ∈ t → a ⊓ b ∈ s ⊼ t := mem_image2_of_mem +lemma infs_subset : s₁ ⊆ s₂ → t₁ ⊆ t₂ → s₁ ⊼ t₁ ⊆ s₂ ⊼ t₂ := image2_subset +lemma infs_subset_left : t₁ ⊆ t₂ → s ⊼ t₁ ⊆ s ⊼ t₂ := image2_subset_left +lemma infs_subset_right : s₁ ⊆ s₂ → s₁ ⊼ t ⊆ s₂ ⊼ t := image2_subset_right + +lemma image_subset_infs_left : b ∈ t → (λ a, a ⊓ b) '' s ⊆ s ⊼ t := image_subset_image2_left +lemma image_subset_infs_right : a ∈ s → (⊓) a '' t ⊆ s ⊼ t := image_subset_image2_right + +lemma forall_infs_iff {p : α → Prop} : (∀ c ∈ s ⊼ t, p c) ↔ ∀ (a ∈ s) (b ∈ t), p (a ⊓ b) := +forall_image2_iff + +@[simp] lemma infs_subset_iff : s ⊼ t ⊆ u ↔ ∀ (a ∈ s) (b ∈ t), a ⊓ b ∈ u := image2_subset_iff + +@[simp] lemma infs_nonempty : (s ⊼ t).nonempty ↔ s.nonempty ∧ t.nonempty := image2_nonempty_iff + +protected lemma nonempty.infs : s.nonempty → t.nonempty → (s ⊼ t).nonempty := nonempty.image2 +lemma nonempty.of_infs_left : (s ⊼ t).nonempty → s.nonempty := nonempty.of_image2_left +lemma nonempty.of_infs_right : (s ⊼ t).nonempty → t.nonempty := nonempty.of_image2_right + +@[simp] lemma empty_infs : ∅ ⊼ t = ∅ := image2_empty_left +@[simp] lemma infs_empty : s ⊼ ∅ = ∅ := image2_empty_right +@[simp] lemma infs_eq_empty : s ⊼ t = ∅ ↔ s = ∅ ∨ t = ∅ := image2_eq_empty_iff + +@[simp] lemma singleton_infs : {a} ⊼ t = t.image (λ b, a ⊓ b) := image2_singleton_left +@[simp] lemma infs_singleton : s ⊼ {b} = s.image (λ a, a ⊓ b) := image2_singleton_right +lemma singleton_infs_singleton : ({a} ⊼ {b} : set α) = {a ⊓ b} := image2_singleton + +lemma infs_union_left : (s₁ ∪ s₂) ⊼ t = s₁ ⊼ t ∪ s₂ ⊼ t := image2_union_left +lemma infs_union_right : s ⊼ (t₁ ∪ t₂) = s ⊼ t₁ ∪ s ⊼ t₂ := image2_union_right + +lemma infs_inter_subset_left : (s₁ ∩ s₂) ⊼ t ⊆ s₁ ⊼ t ∩ s₂ ⊼ t := image2_inter_subset_left +lemma infs_inter_subset_right : s ⊼ (t₁ ∩ t₂) ⊆ s ⊼ t₁ ∩ s ⊼ t₂ := image2_inter_subset_right + +variables (s t u v) + +lemma Union_image_inf_left : (⋃ a ∈ s, (⊓) a '' t) = s ⊼ t := Union_image_left _ +lemma Union_image_inf_right : (⋃ b ∈ t, (⊓ b) '' s) = s ⊼ t := Union_image_right _ + +@[simp] lemma image_inf_prod (s t : set α) : (s ×ˢ t).image (uncurry (⊓)) = s ⊼ t := +image_uncurry_prod _ _ _ + +lemma infs_assoc : (s ⊼ t) ⊼ u = s ⊼ (t ⊼ u) := image2_assoc $ λ _ _ _, inf_assoc +lemma infs_comm : s ⊼ t = t ⊼ s := image2_comm $ λ _ _, inf_comm +lemma infs_left_comm : s ⊼ (t ⊼ u) = t ⊼ (s ⊼ u) := image2_left_comm inf_left_comm +lemma infs_right_comm : (s ⊼ t) ⊼ u = (s ⊼ u) ⊼ t := image2_right_comm inf_right_comm +lemma infs_infs_infs_comm : (s ⊼ t) ⊼ (u ⊼ v) = (s ⊼ u) ⊼ (t ⊼ v) := +image2_image2_image2_comm inf_inf_inf_comm + +end infs + +open_locale set_family + +section distrib_lattice +variables [distrib_lattice α] (s t u : set α) + +lemma sups_infs_subset_left : s ⊻ (t ⊼ u) ⊆ (s ⊻ t) ⊼ (s ⊻ u) := +image2_distrib_subset_left $ λ _ _ _, sup_inf_left + +lemma sups_infs_subset_right : (t ⊼ u) ⊻ s ⊆ (t ⊻ s) ⊼ (u ⊻ s) := +image2_distrib_subset_right $ λ _ _ _, sup_inf_right + +lemma infs_sups_subset_left : s ⊼ (t ⊻ u) ⊆ (s ⊼ t) ⊻ (s ⊼ u) := +image2_distrib_subset_left $ λ _ _ _, inf_sup_left + +lemma infs_sups_subset_right : (t ⊻ u) ⊼ s ⊆ (t ⊼ s) ⊻ (u ⊼ s) := +image2_distrib_subset_right $ λ _ _ _, inf_sup_right + +end distrib_lattice + +end set + +open_locale set_family + +@[simp] lemma upper_closure_sups [semilattice_sup α] (s t : set α) : + upper_closure (s ⊻ t) = upper_closure s ⊔ upper_closure t := +begin + ext a, + simp only [set_like.mem_coe, mem_upper_closure, set.mem_sups, exists_and_distrib_left, + exists_prop, upper_set.coe_sup, set.mem_inter_iff], + split, + { rintro ⟨_, ⟨b, hb, c, hc, rfl⟩, ha⟩, + exact ⟨⟨b, hb, le_sup_left.trans ha⟩, c, hc, le_sup_right.trans ha⟩ }, + { rintro ⟨⟨b, hb, hab⟩, c, hc, hac⟩, + exact ⟨_, ⟨b, hb, c, hc, rfl⟩, sup_le hab hac⟩ } +end + +@[simp] lemma lower_closure_infs [semilattice_inf α] (s t : set α) : + lower_closure (s ⊼ t) = lower_closure s ⊓ lower_closure t := +begin + ext a, + simp only [set_like.mem_coe, mem_lower_closure, set.mem_infs, exists_and_distrib_left, + exists_prop, lower_set.coe_sup, set.mem_inter_iff], + split, + { rintro ⟨_, ⟨b, hb, c, hc, rfl⟩, ha⟩, + exact ⟨⟨b, hb, ha.trans inf_le_left⟩, c, hc, ha.trans inf_le_right⟩ }, + { rintro ⟨⟨b, hb, hab⟩, c, hc, hac⟩, + exact ⟨_, ⟨b, hb, c, hc, rfl⟩, le_inf hab hac⟩ } +end From e085d1df33274f4b32f611f483aae678ba0b42df Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sat, 18 Mar 2023 04:16:55 +0000 Subject: [PATCH 0653/1291] feat(algebra/quadratic_discriminant): generalize, use `ne_zero` (#18606) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit * Add `discrim_neg`. * Add `discrim_eq_sq_of_quadratic_eq_zero`, an implication from `quadratic_eq_zero_iff_discrim_eq_sq` that doesn't need extra assumptions. * Assume `[ne_zero (2 : _)]` instead of `2 ≠ 0` or `[invertible (2 : _)]`. * Drop unneeded assumptions in `quadratic_ne_zero_of_discrim_ne_sq`, use `s ^ 2` instead of `s * s`. * Add `discrim_le_zero_of_nonpos` and `discrim_lt_zero_of_neg`. --- src/algebra/quadratic_discriminant.lean | 70 ++++++++++++++----------- 1 file changed, 39 insertions(+), 31 deletions(-) diff --git a/src/algebra/quadratic_discriminant.lean b/src/algebra/quadratic_discriminant.lean index ea3439b5e939a..2821a3d9efe3e 100644 --- a/src/algebra/quadratic_discriminant.lean +++ b/src/algebra/quadratic_discriminant.lean @@ -25,6 +25,8 @@ This file defines the discriminant of a quadratic and gives the solution to a qu - `quadratic_ne_zero_of_discrim_ne_sq`: if the discriminant has no square root, then the corresponding quadratic has no root. - `discrim_le_zero`: if a quadratic is always non-negative, then its discriminant is non-positive. +- `discrim_le_zero_of_nonpos`, `discrim_lt_zero`, `discrim_lt_zero_of_neg`: versions of this + statement with other inequalities. ## Tags @@ -39,46 +41,47 @@ variables {R : Type*} /-- Discriminant of a quadratic -/ def discrim [ring R] (a b c : R) : R := b^2 - 4 * a * c -variables [comm_ring R] [is_domain R] {a b c : R} +@[simp] lemma discrim_neg [ring R] (a b c : R) : discrim (-a) (-b) (-c) = discrim a b c := +by simp [discrim] + +variables [comm_ring R] {a b c : R} + +lemma discrim_eq_sq_of_quadratic_eq_zero {x : R} (h : a * x * x + b * x + c = 0) : + discrim a b c = (2 * a * x + b) ^ 2 := +begin + rw [discrim], + linear_combination -4 * a * h +end /-- A quadratic has roots if and only if its discriminant equals some square. -/ -lemma quadratic_eq_zero_iff_discrim_eq_sq (h2 : (2 : R) ≠ 0) (ha : a ≠ 0) (x : R) : +lemma quadratic_eq_zero_iff_discrim_eq_sq [ne_zero (2 : R)] [no_zero_divisors R] + (ha : a ≠ 0) {x : R} : a * x * x + b * x + c = 0 ↔ discrim a b c = (2 * a * x + b) ^ 2 := begin - dsimp [discrim] at *, - split, - { assume h, - linear_combination -4 * a * h }, - { assume h, - have ha : 2 * 2 * a ≠ 0 := mul_ne_zero (mul_ne_zero h2 h2) ha, - apply mul_left_cancel₀ ha, - linear_combination -h } + refine ⟨discrim_eq_sq_of_quadratic_eq_zero, λ h, _⟩, + rw [discrim] at h, + have ha : 2 * 2 * a ≠ 0 := mul_ne_zero (mul_ne_zero (ne_zero.ne _) (ne_zero.ne _)) ha, + apply mul_left_cancel₀ ha, + linear_combination -h end /-- A quadratic has no root if its discriminant has no square root. -/ -lemma quadratic_ne_zero_of_discrim_ne_sq (h2 : (2 : R) ≠ 0) (ha : a ≠ 0) - (h : ∀ s : R, discrim a b c ≠ s * s) (x : R) : +lemma quadratic_ne_zero_of_discrim_ne_sq (h : ∀ s : R, discrim a b c ≠ s^2) (x : R) : a * x * x + b * x + c ≠ 0 := -begin - assume h', - rw [quadratic_eq_zero_iff_discrim_eq_sq h2 ha, sq] at h', - exact h _ h' -end +mt discrim_eq_sq_of_quadratic_eq_zero $ h _ end ring section field -variables {K : Type*} [field K] [invertible (2 : K)] {a b c x : K} +variables {K : Type*} [field K] [ne_zero (2 : K)] {a b c x : K} -/-- Roots of a quadratic -/ +/-- Roots of a quadratic equation. -/ lemma quadratic_eq_zero_iff (ha : a ≠ 0) {s : K} (h : discrim a b c = s * s) (x : K) : a * x * x + b * x + c = 0 ↔ x = (-b + s) / (2 * a) ∨ x = (-b - s) / (2 * a) := begin - have h2 : (2 : K) ≠ 0 := nonzero_of_invertible 2, - rw [quadratic_eq_zero_iff_discrim_eq_sq h2 ha, h, sq, mul_self_eq_mul_self_iff], - have ne : 2 * a ≠ 0 := mul_ne_zero h2 ha, + rw [quadratic_eq_zero_iff_discrim_eq_sq ha, h, sq, mul_self_eq_mul_self_iff], field_simp, apply or_congr, { split; intro h'; linear_combination -h' }, @@ -108,7 +111,7 @@ end field section linear_ordered_field variables {K : Type*} [linear_ordered_field K] {a b c : K} -/-- If a polynomial of degree 2 is always nonnegative, then its discriminant is nonpositive -/ +/-- If a polynomial of degree 2 is always nonnegative, then its discriminant is nonpositive. -/ lemma discrim_le_zero (h : ∀ x : K, 0 ≤ a * x * x + b * x + c) : discrim a b c ≤ 0 := begin rw [discrim, sq], @@ -120,19 +123,20 @@ begin rcases (this.eventually (eventually_lt_at_bot 0)).exists with ⟨x, hx⟩, exact false.elim ((h x).not_lt $ by rwa ← add_mul) }, -- if a = 0 - { rcases em (b = 0) with (rfl|hb), + { rcases eq_or_ne b 0 with (rfl|hb), { simp }, { have := h ((-c - 1) / b), rw [mul_div_cancel' _ hb] at this, linarith } }, -- if a > 0 - { have := calc - 4 * a * (a * (-(b / a) * (1 / 2)) * (-(b / a) * (1 / 2)) + b * (-(b / a) * (1 / 2)) + c) - = (a * (b / a)) * (a * (b / a)) - 2 * (a * (b / a)) * b + 4 * a * c : by ring - ... = -(b * b - 4 * a * c) : by { simp only [mul_div_cancel' b (ne_of_gt ha)], ring }, - have ha' : 0 ≤ 4 * a, by linarith, - have h := (mul_nonneg ha' (h (-(b / a) * (1 / 2)))), - rw this at h, rwa ← neg_nonneg } + { have ha' : 0 ≤ 4 * a := mul_nonneg zero_le_four ha.le, + have := h (-b / (2 * a)), + convert neg_nonpos.2 (mul_nonneg ha' (h (-b / (2 * a)))), + field_simp [ha.ne'], + ring } end +lemma discrim_le_zero_of_nonpos (h : ∀ x : K, a * x * x + b * x + c ≤ 0) : discrim a b c ≤ 0 := +discrim_neg a b c ▸ discrim_le_zero (by simpa only [neg_mul, ← neg_add, neg_nonneg]) + /-- If a polynomial of degree 2 is always positive, then its discriminant is negative, at least when the coefficient of the quadratic term is nonzero. @@ -148,4 +152,8 @@ begin linarith end +lemma discrim_lt_zero_of_neg (ha : a ≠ 0) (h : ∀ x : K, a * x * x + b * x + c < 0) : + discrim a b c < 0 := +discrim_neg a b c ▸ discrim_lt_zero (neg_ne_zero.2 ha) (by simpa only [neg_mul, ← neg_add, neg_pos]) + end linear_ordered_field From 932872382355f00112641d305ba0619305dc8642 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Sat, 18 Mar 2023 06:40:16 +0000 Subject: [PATCH 0654/1291] chore(*): add mathlib4 synchronization comments (#18595) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.algebra.subalgebra.pointwise` * `algebra.algebra.unitization` * `algebra.char_p.exp_char` * `algebra.direct_sum.module` * `algebra.field.ulift` * `algebra.monoid_algebra.degree` * `algebra.monoid_algebra.no_zero_divisors` * `algebra.triv_sq_zero_ext` * `category_theory.category.Bipointed` * `category_theory.concrete_category.bundled_hom` * `category_theory.elementwise` * `category_theory.preadditive.of_biproducts` * `data.complex.exponential` * `data.mv_polynomial.basic` * `data.mv_polynomial.comap` * `data.mv_polynomial.counit` * `data.mv_polynomial.rename` * `data.polynomial.cancel_leads` * `data.polynomial.identities` * `group_theory.perm.cycle.basic` * `measure_theory.card_measurable_space` * `ring_theory.algebra_tower` * `ring_theory.coprime.ideal` * `ring_theory.free_ring` * `ring_theory.ideal.prod` * `ring_theory.mv_polynomial.tower` * `ring_theory.polynomial.opposites` * `ring_theory.subring.pointwise` * `ring_theory.valuation.basic` * `topology.algebra.nonarchimedean.basic` * `topology.algebra.open_subgroup` * `topology.algebra.with_zero_topology` * `topology.homotopy.equiv` --- src/algebra/algebra/subalgebra/pointwise.lean | 3 +++ src/algebra/algebra/unitization.lean | 3 +++ src/algebra/char_p/exp_char.lean | 3 +++ src/algebra/direct_sum/module.lean | 3 +++ src/algebra/field/ulift.lean | 3 +++ src/algebra/monoid_algebra/degree.lean | 3 +++ src/algebra/monoid_algebra/no_zero_divisors.lean | 3 +++ src/algebra/triv_sq_zero_ext.lean | 3 +++ src/category_theory/category/Bipointed.lean | 3 +++ src/category_theory/concrete_category/bundled_hom.lean | 3 +++ src/category_theory/elementwise.lean | 3 +++ src/category_theory/preadditive/of_biproducts.lean | 3 +++ src/data/complex/exponential.lean | 3 +++ src/data/mv_polynomial/basic.lean | 3 +++ src/data/mv_polynomial/comap.lean | 3 +++ src/data/mv_polynomial/counit.lean | 3 +++ src/data/mv_polynomial/rename.lean | 3 +++ src/data/polynomial/cancel_leads.lean | 3 +++ src/data/polynomial/identities.lean | 3 +++ src/group_theory/perm/cycle/basic.lean | 3 +++ src/measure_theory/card_measurable_space.lean | 3 +++ src/ring_theory/algebra_tower.lean | 3 +++ src/ring_theory/coprime/ideal.lean | 3 +++ src/ring_theory/free_ring.lean | 3 +++ src/ring_theory/ideal/prod.lean | 3 +++ src/ring_theory/mv_polynomial/tower.lean | 3 +++ src/ring_theory/polynomial/opposites.lean | 3 +++ src/ring_theory/subring/pointwise.lean | 3 +++ src/ring_theory/valuation/basic.lean | 3 +++ src/topology/algebra/nonarchimedean/basic.lean | 3 +++ src/topology/algebra/open_subgroup.lean | 3 +++ src/topology/algebra/with_zero_topology.lean | 3 +++ src/topology/homotopy/equiv.lean | 3 +++ 33 files changed, 99 insertions(+) diff --git a/src/algebra/algebra/subalgebra/pointwise.lean b/src/algebra/algebra/subalgebra/pointwise.lean index 70ffd9895815a..a4183947fa4f7 100644 --- a/src/algebra/algebra/subalgebra/pointwise.lean +++ b/src/algebra/algebra/subalgebra/pointwise.lean @@ -11,6 +11,9 @@ import ring_theory.adjoin.basic /-! # Pointwise actions on subalgebras. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + If `R'` acts on an `R`-algebra `A` (so that `R'` and `R` actions commute) then we get an `R'` action on the collection of `R`-subalgebras. -/ diff --git a/src/algebra/algebra/unitization.lean b/src/algebra/algebra/unitization.lean index ce86cba306cb8..a789b95747b4b 100644 --- a/src/algebra/algebra/unitization.lean +++ b/src/algebra/algebra/unitization.lean @@ -11,6 +11,9 @@ import algebra.hom.non_unital_alg /-! # Unitization of a non-unital algebra +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given a non-unital `R`-algebra `A` (given via the type classes `[non_unital_ring A] [module R A] [smul_comm_class R A A] [is_scalar_tower R A A]`) we construct the minimal unital `R`-algebra containing `A` as an ideal. This object `algebra.unitization R A` is diff --git a/src/algebra/char_p/exp_char.lean b/src/algebra/char_p/exp_char.lean index 4aeb0a2af172e..86273305cea30 100644 --- a/src/algebra/char_p/exp_char.lean +++ b/src/algebra/char_p/exp_char.lean @@ -9,6 +9,9 @@ import data.nat.prime /-! # Exponential characteristic +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the exponential characteristic and establishes a few basic results relating it to the (ordinary characteristic). The definition is stated for a semiring, but the actual results are for nontrivial rings diff --git a/src/algebra/direct_sum/module.lean b/src/algebra/direct_sum/module.lean index 560892bc2a36d..6735c5d53901d 100644 --- a/src/algebra/direct_sum/module.lean +++ b/src/algebra/direct_sum/module.lean @@ -9,6 +9,9 @@ import linear_algebra.dfinsupp /-! # Direct sum of modules +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The first part of the file provides constructors for direct sums of modules. It provides a construction of the direct sum using the universal property and proves its uniqueness (`direct_sum.to_module.unique`). diff --git a/src/algebra/field/ulift.lean b/src/algebra/field/ulift.lean index fe49311689450..a515210095d80 100644 --- a/src/algebra/field/ulift.lean +++ b/src/algebra/field/ulift.lean @@ -9,6 +9,9 @@ import algebra.ring.ulift /-! # Field instances for `ulift` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines instances for field, semifield and related structures on `ulift` types. (Recall `ulift α` is just a "copy" of a type `α` in a higher universe.) diff --git a/src/algebra/monoid_algebra/degree.lean b/src/algebra/monoid_algebra/degree.lean index 602c3a8fa079f..2aae4a028c4e4 100644 --- a/src/algebra/monoid_algebra/degree.lean +++ b/src/algebra/monoid_algebra/degree.lean @@ -8,6 +8,9 @@ import algebra.monoid_algebra.support /-! # Lemmas about the `sup` and `inf` of the support of `add_monoid_algebra` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## TODO The current plan is to state and prove lemmas about `finset.sup (finsupp.support f) D` with a "generic" degree/weight function `D` from the grading Type `A` to a somewhat ordered Type `B`. diff --git a/src/algebra/monoid_algebra/no_zero_divisors.lean b/src/algebra/monoid_algebra/no_zero_divisors.lean index 4a6499cd4e067..58493153c3cc5 100644 --- a/src/algebra/monoid_algebra/no_zero_divisors.lean +++ b/src/algebra/monoid_algebra/no_zero_divisors.lean @@ -8,6 +8,9 @@ import algebra.monoid_algebra.support /-! # Variations on non-zero divisors in `add_monoid_algebra`s +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file studies the interaction between typeclass assumptions on two Types `R` and `A` and whether `add_monoid_algebra R A` has non-zero zero-divisors. For some background on related questions, see [Kaplansky's Conjectures](https://en.wikipedia.org/wiki/Kaplansky%27s_conjectures), diff --git a/src/algebra/triv_sq_zero_ext.lean b/src/algebra/triv_sq_zero_ext.lean index fc2bedc4c9c16..f2ffc13cfe006 100644 --- a/src/algebra/triv_sq_zero_ext.lean +++ b/src/algebra/triv_sq_zero_ext.lean @@ -10,6 +10,9 @@ import linear_algebra.prod /-! # Trivial Square-Zero Extension +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given a ring `R` together with an `(R, R)`-bimodule `M`, the trivial square-zero extension of `M` over `R` is defined to be the `R`-algebra `R ⊕ M` with multiplication given by `(r₁ + m₁) * (r₂ + m₂) = r₁ r₂ + r₁ m₂ + m₁ r₂`. diff --git a/src/category_theory/category/Bipointed.lean b/src/category_theory/category/Bipointed.lean index bafe7996568ea..12a348f069d99 100644 --- a/src/category_theory/category/Bipointed.lean +++ b/src/category_theory/category/Bipointed.lean @@ -8,6 +8,9 @@ import category_theory.category.Pointed /-! # The category of bipointed types +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This defines `Bipointed`, the category of bipointed types. ## TODO diff --git a/src/category_theory/concrete_category/bundled_hom.lean b/src/category_theory/concrete_category/bundled_hom.lean index 0791c841b0509..7676d31b2ea94 100644 --- a/src/category_theory/concrete_category/bundled_hom.lean +++ b/src/category_theory/concrete_category/bundled_hom.lean @@ -9,6 +9,9 @@ import category_theory.concrete_category.bundled /-! # Category instances for algebraic structures that use bundled homs. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Many algebraic structures in Lean initially used unbundled homs (e.g. a bare function between types, along with an `is_monoid_hom` typeclass), but the general trend is towards using bundled homs. diff --git a/src/category_theory/elementwise.lean b/src/category_theory/elementwise.lean index 8be2fd34d558a..e88f0387888e4 100644 --- a/src/category_theory/elementwise.lean +++ b/src/category_theory/elementwise.lean @@ -9,6 +9,9 @@ import category_theory.concrete_category.basic /-! # Use the `elementwise` attribute to create applied versions of lemmas. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Usually we would use `@[elementwise]` at the point of definition, however some early parts of the category theory library are imported by `tactic.elementwise`, so we need to add the attribute after the fact. diff --git a/src/category_theory/preadditive/of_biproducts.lean b/src/category_theory/preadditive/of_biproducts.lean index d1f0cc0508307..6deac1483f48f 100644 --- a/src/category_theory/preadditive/of_biproducts.lean +++ b/src/category_theory/preadditive/of_biproducts.lean @@ -9,6 +9,9 @@ import group_theory.eckmann_hilton /-! # Constructing a semiadditive structure from binary biproducts +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We show that any category with zero morphisms and binary biproducts is enriched over the category of commutative monoids. diff --git a/src/data/complex/exponential.lean b/src/data/complex/exponential.lean index 4aebd98bd083d..f7a9a272be1af 100644 --- a/src/data/complex/exponential.lean +++ b/src/data/complex/exponential.lean @@ -10,6 +10,9 @@ import data.nat.choose.sum /-! # Exponential, trigonometric and hyperbolic trigonometric functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains the definitions of the real and complex exponential, sine, cosine, tangent, hyperbolic sine, hyperbolic cosine, and hyperbolic tangent functions. diff --git a/src/data/mv_polynomial/basic.lean b/src/data/mv_polynomial/basic.lean index 4d4fefcd20717..973a34c4ca479 100644 --- a/src/data/mv_polynomial/basic.lean +++ b/src/data/mv_polynomial/basic.lean @@ -13,6 +13,9 @@ import ring_theory.adjoin.basic /-! # Multivariate polynomials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines polynomial rings over a base ring (or even semiring), with variables from a general type `σ` (which could be infinite). diff --git a/src/data/mv_polynomial/comap.lean b/src/data/mv_polynomial/comap.lean index a532cd8918d99..7dcbdb07d6f48 100644 --- a/src/data/mv_polynomial/comap.lean +++ b/src/data/mv_polynomial/comap.lean @@ -9,6 +9,9 @@ import data.mv_polynomial.rename /-! # `comap` operation on `mv_polynomial` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the `comap` function on `mv_polynomial`. `mv_polynomial.comap` is a low-tech example of a map of "algebraic varieties," modulo the fact that diff --git a/src/data/mv_polynomial/counit.lean b/src/data/mv_polynomial/counit.lean index 8488c8ff2b482..85b80de52a887 100644 --- a/src/data/mv_polynomial/counit.lean +++ b/src/data/mv_polynomial/counit.lean @@ -9,6 +9,9 @@ import data.mv_polynomial.basic /-! ## Counit morphisms for multivariate polynomials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + One may consider the ring of multivariate polynomials `mv_polynomial A R` with coefficients in `R` and variables indexed by `A`. If `A` is not just a type, but an algebra over `R`, then there is a natural surjective algebra homomorphism `mv_polynomial A R →ₐ[R] A` diff --git a/src/data/mv_polynomial/rename.lean b/src/data/mv_polynomial/rename.lean index d4239c2e5f58b..858e5b2cf69d0 100644 --- a/src/data/mv_polynomial/rename.lean +++ b/src/data/mv_polynomial/rename.lean @@ -9,6 +9,9 @@ import data.mv_polynomial.basic /-! # Renaming variables of polynomials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file establishes the `rename` operation on multivariate polynomials, which modifies the set of variables. diff --git a/src/data/polynomial/cancel_leads.lean b/src/data/polynomial/cancel_leads.lean index 2592e08f3b9ce..4275643fba7da 100644 --- a/src/data/polynomial/cancel_leads.lean +++ b/src/data/polynomial/cancel_leads.lean @@ -9,6 +9,9 @@ import data.polynomial.degree.lemmas /-! # Cancel the leading terms of two polynomials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Definition * `cancel_leads p q`: the polynomial formed by multiplying `p` and `q` by monomials so that they diff --git a/src/data/polynomial/identities.lean b/src/data/polynomial/identities.lean index 642cf06533b7d..93d50ba61afbe 100644 --- a/src/data/polynomial/identities.lean +++ b/src/data/polynomial/identities.lean @@ -10,6 +10,9 @@ import tactic.ring_exp /-! # Theory of univariate polynomials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The main def is `binom_expansion`. -/ diff --git a/src/group_theory/perm/cycle/basic.lean b/src/group_theory/perm/cycle/basic.lean index c5d400b19a151..24d92b6a295be 100644 --- a/src/group_theory/perm/cycle/basic.lean +++ b/src/group_theory/perm/cycle/basic.lean @@ -13,6 +13,9 @@ import logic.equiv.fintype /-! # Cyclic permutations +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file develops the theory of cycles in permutations. ## Main definitions diff --git a/src/measure_theory/card_measurable_space.lean b/src/measure_theory/card_measurable_space.lean index 33bda61263722..d7dfb8c03f691 100644 --- a/src/measure_theory/card_measurable_space.lean +++ b/src/measure_theory/card_measurable_space.lean @@ -10,6 +10,9 @@ import set_theory.cardinal.continuum /-! # Cardinal of sigma-algebras +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + If a sigma-algebra is generated by a set of sets `s`, then the cardinality of the sigma-algebra is bounded by `(max (#s) 2) ^ ℵ₀`. This is stated in `measurable_space.cardinal_generate_measurable_le` and `measurable_space.cardinal_measurable_set_le`. diff --git a/src/ring_theory/algebra_tower.lean b/src/ring_theory/algebra_tower.lean index 85128910d2ca4..84b1b34d5c37c 100644 --- a/src/ring_theory/algebra_tower.lean +++ b/src/ring_theory/algebra_tower.lean @@ -11,6 +11,9 @@ import linear_algebra.basis /-! # Towers of algebras +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We set up the basic theory of algebra towers. An algebra tower A/S/R is expressed by having instances of `algebra A S`, `algebra R S`, `algebra R A` and `is_scalar_tower R S A`, the later asserting the diff --git a/src/ring_theory/coprime/ideal.lean b/src/ring_theory/coprime/ideal.lean index 9e27810bf08ca..1cf3d3d97e7c0 100644 --- a/src/ring_theory/coprime/ideal.lean +++ b/src/ring_theory/coprime/ideal.lean @@ -9,6 +9,9 @@ import ring_theory.ideal.operations /-! # An additional lemma about coprime ideals +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This lemma generalises `exists_sum_eq_one_iff_pairwise_coprime` to the case of non-principal ideals. It is on a separate file due to import requirements. -/ diff --git a/src/ring_theory/free_ring.lean b/src/ring_theory/free_ring.lean index 026bd954c7a20..983f8c47dac84 100644 --- a/src/ring_theory/free_ring.lean +++ b/src/ring_theory/free_ring.lean @@ -8,6 +8,9 @@ import group_theory.free_abelian_group /-! # Free rings +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The theory of the free ring over a type. ## Main definitions diff --git a/src/ring_theory/ideal/prod.lean b/src/ring_theory/ideal/prod.lean index 081fdbb7af8fb..96485fbc8535e 100644 --- a/src/ring_theory/ideal/prod.lean +++ b/src/ring_theory/ideal/prod.lean @@ -8,6 +8,9 @@ import ring_theory.ideal.operations /-! # Ideals in product rings +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + For commutative rings `R` and `S` and ideals `I ≤ R`, `J ≤ S`, we define `ideal.prod I J` as the product `I × J`, viewed as an ideal of `R × S`. In `ideal_prod_eq` we show that every ideal of `R × S` is of this form. Furthermore, we show that every prime ideal of `R × S` is of the form diff --git a/src/ring_theory/mv_polynomial/tower.lean b/src/ring_theory/mv_polynomial/tower.lean index 9877c4389812b..ddfee025e7909 100644 --- a/src/ring_theory/mv_polynomial/tower.lean +++ b/src/ring_theory/mv_polynomial/tower.lean @@ -10,6 +10,9 @@ import data.mv_polynomial.basic /-! # Algebra towers for multivariate polynomial +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves some basic results about the algebra tower structure for the type `mv_polynomial σ R`. diff --git a/src/ring_theory/polynomial/opposites.lean b/src/ring_theory/polynomial/opposites.lean index 7bb6bcc73ff35..4ecea9805d529 100644 --- a/src/ring_theory/polynomial/opposites.lean +++ b/src/ring_theory/polynomial/opposites.lean @@ -8,6 +8,9 @@ import data.polynomial.degree.definitions /-! # Interactions between `R[X]` and `Rᵐᵒᵖ[X]` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains the basic API for "pushing through" the isomorphism `op_ring_equiv : R[X]ᵐᵒᵖ ≃+* Rᵐᵒᵖ[X]`. It allows going back and forth between a polynomial ring over a semiring and the polynomial ring over the opposite semiring. -/ diff --git a/src/ring_theory/subring/pointwise.lean b/src/ring_theory/subring/pointwise.lean index f5882efd1fd8f..86242c30be28e 100644 --- a/src/ring_theory/subring/pointwise.lean +++ b/src/ring_theory/subring/pointwise.lean @@ -10,6 +10,9 @@ import data.set.pointwise.basic /-! # Pointwise instances on `subring`s +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file provides the action `subring.pointwise_mul_action` which matches the action of `mul_action_set`. diff --git a/src/ring_theory/valuation/basic.lean b/src/ring_theory/valuation/basic.lean index 6acdabe47103b..0f209fc8b7965 100644 --- a/src/ring_theory/valuation/basic.lean +++ b/src/ring_theory/valuation/basic.lean @@ -11,6 +11,9 @@ import ring_theory.ideal.operations # The basics of valuation theory. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The basic theory of valuations (non-archimedean norms) on a commutative ring, following T. Wedhorn's unpublished notes “Adic Spaces” ([wedhorn_adic]). diff --git a/src/topology/algebra/nonarchimedean/basic.lean b/src/topology/algebra/nonarchimedean/basic.lean index a19b98853f8de..8f6d94f28a25a 100644 --- a/src/topology/algebra/nonarchimedean/basic.lean +++ b/src/topology/algebra/nonarchimedean/basic.lean @@ -10,6 +10,9 @@ import topology.algebra.ring.basic /-! # Nonarchimedean Topology +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we set up the theory of nonarchimedean topological groups and rings. A nonarchimedean group is a topological group whose topology admits a basis of diff --git a/src/topology/algebra/open_subgroup.lean b/src/topology/algebra/open_subgroup.lean index 6f08dd30e8435..35d5a9753cca5 100644 --- a/src/topology/algebra/open_subgroup.lean +++ b/src/topology/algebra/open_subgroup.lean @@ -9,6 +9,9 @@ import topology.sets.opens /-! # Open subgroups of a topological groups +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This files builds the lattice `open_subgroup G` of open subgroups in a topological group `G`, and its additive version `open_add_subgroup`. This lattice has a top element, the subgroup of all elements, but no bottom element in general. The trivial subgroup which is the natural candidate diff --git a/src/topology/algebra/with_zero_topology.lean b/src/topology/algebra/with_zero_topology.lean index 4332d27a9da56..4c793959b3aad 100644 --- a/src/topology/algebra/with_zero_topology.lean +++ b/src/topology/algebra/with_zero_topology.lean @@ -9,6 +9,9 @@ import topology.algebra.order.field /-! # The topology on linearly ordered commutative groups with zero +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Let `Γ₀` be a linearly ordered commutative group to which we have adjoined a zero element. Then `Γ₀` may naturally be endowed with a topology that turns `Γ₀` into a topological monoid. Neighborhoods of zero are sets containing `{γ | γ < γ₀}` for some invertible element `γ₀` diff --git a/src/topology/homotopy/equiv.lean b/src/topology/homotopy/equiv.lean index 641495cefeb98..c0542004305c3 100644 --- a/src/topology/homotopy/equiv.lean +++ b/src/topology/homotopy/equiv.lean @@ -10,6 +10,9 @@ import topology.homotopy.basic # Homotopy equivalences between topological spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we define homotopy equivalences between topological spaces `X` and `Y` as a pair of functions `f : C(X, Y)` and `g : C(Y, X)` such that `f.comp g` and `g.comp f` are both homotopic to `id`. From 027ff1e4f0b337893269a7dcd90476ceab95570b Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sat, 18 Mar 2023 09:58:11 +0000 Subject: [PATCH 0655/1291] chore(geometry/euclidean/sphere/basic): generalize typeclasses (#18605) `cospherical_empty` asks for `P` to be a real inner product space when all it needs is `nonempty P`. The motivation for this PR is that it fixes a linter error in #18583; lean complains that we are carrying around the type `V` for no reason. The alternative here would be to force all these typeclass arguments to be present in the constructor for `euclidean_space.sphere` and use `no_lint unused_arguments` --- src/geometry/euclidean/sphere/basic.lean | 90 ++++++++++++------------ 1 file changed, 45 insertions(+), 45 deletions(-) diff --git a/src/geometry/euclidean/sphere/basic.lean b/src/geometry/euclidean/sphere/basic.lean index 6eaf3054a2cea..91372a3c373a7 100644 --- a/src/geometry/euclidean/sphere/basic.lean +++ b/src/geometry/euclidean/sphere/basic.lean @@ -29,18 +29,21 @@ open_locale real_inner_product_space namespace euclidean_geometry -variables {V : Type*} (P : Type*) [inner_product_space ℝ V] [metric_space P] [normed_add_torsor V P] +variables {V : Type*} (P : Type*) open finite_dimensional /-- A `sphere P` bundles a `center` and `radius`. This definition does not require the radius to be positive; that should be given as a hypothesis to lemmas that require it. -/ -@[ext] structure sphere := +@[ext] structure sphere [metric_space P] := (center : P) (radius : ℝ) variables {P} +section metric_space +variables [metric_space P] + instance [nonempty P] : nonempty (sphere P) := ⟨⟨classical.arbitrary P, 0⟩⟩ instance : has_coe (sphere P) (set P) := ⟨λ s, metric.sphere s.center s.radius⟩ @@ -132,16 +135,9 @@ begin exact ⟨c, r, λ p hp, hcr p (hs hp)⟩ end -include V - /-- The empty set is cospherical. -/ -lemma cospherical_empty : cospherical (∅ : set P) := -begin - use add_torsor.nonempty.some, - simp, -end - -omit V +lemma cospherical_empty [nonempty P] : cospherical (∅ : set P) := +let ⟨p⟩ := ‹nonempty P› in ⟨p, 0, λ p, false.elim⟩ /-- A single point is cospherical. -/ lemma cospherical_singleton (p : P) : cospherical ({p} : set P) := @@ -150,23 +146,47 @@ begin simp end +end metric_space + +section normed_space +variables [normed_add_comm_group V] [normed_space ℝ V] [metric_space P] [normed_add_torsor V P] include V /-- Two points are cospherical. -/ lemma cospherical_pair (p₁ p₂ : P) : cospherical ({p₁, p₂} : set P) := -begin - use [(2⁻¹ : ℝ) • (p₂ -ᵥ p₁) +ᵥ p₁, (2⁻¹ : ℝ) * (dist p₂ p₁)], - intro p, - rw [set.mem_insert_iff, set.mem_singleton_iff], - rintro ⟨_|_⟩, - { rw [dist_eq_norm_vsub V p₁, vsub_vadd_eq_vsub_sub, vsub_self, zero_sub, norm_neg, norm_smul, - dist_eq_norm_vsub V p₂], - simp }, - { rw [H, dist_eq_norm_vsub V p₂, vsub_vadd_eq_vsub_sub, dist_eq_norm_vsub V p₂], - conv_lhs { congr, congr, rw ←one_smul ℝ (p₂ -ᵥ p₁ : V) }, - rw [←sub_smul, norm_smul], - norm_num } -end +⟨midpoint ℝ p₁ p₂, ‖(2 : ℝ)‖⁻¹ * dist p₁ p₂, begin + rintros p (rfl | rfl | _), + { rw [dist_comm, dist_midpoint_left] }, + { rw [dist_comm, dist_midpoint_right] } +end⟩ + +/-- A set of points is concyclic if it is cospherical and coplanar. (Most results are stated +directly in terms of `cospherical` instead of using `concyclic`.) -/ +structure concyclic (ps : set P) : Prop := +(cospherical : cospherical ps) +(coplanar : coplanar ℝ ps) + +/-- A subset of a concyclic set is concyclic. -/ +lemma concyclic.subset {ps₁ ps₂ : set P} (hs : ps₁ ⊆ ps₂) (h : concyclic ps₂) : concyclic ps₁ := +⟨h.1.subset hs, h.2.subset hs⟩ + +/-- The empty set is concyclic. -/ +lemma concyclic_empty : concyclic (∅ : set P) := +⟨cospherical_empty, coplanar_empty ℝ P⟩ + +/-- A single point is concyclic. -/ +lemma concyclic_singleton (p : P) : concyclic ({p} : set P) := +⟨cospherical_singleton p, coplanar_singleton ℝ p⟩ + +/-- Two points are concyclic. -/ +lemma concyclic_pair (p₁ p₂ : P) : concyclic ({p₁, p₂} : set P) := +⟨cospherical_pair p₁ p₂, coplanar_pair ℝ p₁ p₂⟩ + +end normed_space + +section euclidean_space +variables [inner_product_space ℝ V] [metric_space P] [normed_add_torsor V P] +include V /-- Any three points in a cospherical set are affinely independent. -/ lemma cospherical.affine_independent {s : set P} (hs : cospherical s) {p : fin 3 → P} @@ -324,26 +344,6 @@ lemma sbtw_of_collinear_of_dist_center_lt_radius {s : sphere P} {p₁ p₂ p₃ (hp₃ : p₃ ∈ s) (hp₁p₃ : p₁ ≠ p₃) : sbtw ℝ p₁ p₂ p₃ := h.sbtw_of_dist_eq_of_dist_lt hp₁ hp₂ hp₃ hp₁p₃ -/-- A set of points is concyclic if it is cospherical and coplanar. (Most results are stated -directly in terms of `cospherical` instead of using `concyclic`.) -/ -structure concyclic (ps : set P) : Prop := -(cospherical : cospherical ps) -(coplanar : coplanar ℝ ps) - -/-- A subset of a concyclic set is concyclic. -/ -lemma concyclic.subset {ps₁ ps₂ : set P} (hs : ps₁ ⊆ ps₂) (h : concyclic ps₂) : concyclic ps₁ := -⟨h.1.subset hs, h.2.subset hs⟩ - -/-- The empty set is concyclic. -/ -lemma concyclic_empty : concyclic (∅ : set P) := -⟨cospherical_empty, coplanar_empty ℝ P⟩ - -/-- A single point is concyclic. -/ -lemma concyclic_singleton (p : P) : concyclic ({p} : set P) := -⟨cospherical_singleton p, coplanar_singleton ℝ p⟩ - -/-- Two points are concyclic. -/ -lemma concyclic_pair (p₁ p₂ : P) : concyclic ({p₁, p₂} : set P) := -⟨cospherical_pair p₁ p₂, coplanar_pair ℝ p₁ p₂⟩ +end euclidean_space end euclidean_geometry From f430769b562e0cedef59ee1ed968d67e0e0c86ba Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sat, 18 Mar 2023 11:38:18 +0000 Subject: [PATCH 0656/1291] chore(topology/algebra/module/basic): move a lemma to a new file (#18608) Also slightly generalize from `[is_simple_module R R]` to `[is_simple_module R N]`. --- src/topology/algebra/module/basic.lean | 17 +-------- .../algebra/module/finite_dimension.lean | 1 + src/topology/algebra/module/simple.lean | 35 +++++++++++++++++++ 3 files changed, 37 insertions(+), 16 deletions(-) create mode 100644 src/topology/algebra/module/simple.lean diff --git a/src/topology/algebra/module/basic.lean b/src/topology/algebra/module/basic.lean index e6359dc849fd8..6d322735e4f11 100644 --- a/src/topology/algebra/module/basic.lean +++ b/src/topology/algebra/module/basic.lean @@ -12,7 +12,6 @@ import topology.uniform_space.uniform_embedding import algebra.algebra.basic import linear_algebra.projection import linear_algebra.pi -import ring_theory.simple_module /-! # Theory of topological modules and continuous linear maps. @@ -168,13 +167,10 @@ S.to_add_subgroup.topological_add_group end submodule section closure -variables {R R' : Type u} {M M' : Type v} +variables {R : Type u} {M : Type v} [semiring R] [topological_space R] -[ring R'] [topological_space R'] [topological_space M] [add_comm_monoid M] -[topological_space M'] [add_comm_group M'] [module R M] [has_continuous_smul R M] -[module R' M'] [has_continuous_smul R' M'] lemma submodule.closure_smul_self_subset (s : submodule R M) : (λ p : R × M, p.1 • p.2) '' (set.univ ×ˢ closure s) ⊆ closure s := @@ -249,17 +245,6 @@ lemma submodule.is_closed_or_dense_of_is_coatom (s : submodule R M) (hs : is_coa (hs.le_iff.mp s.le_topological_closure).swap.imp (is_closed_of_closure_subset ∘ eq.le) submodule.dense_iff_topological_closure_eq_top.mpr -lemma linear_map.is_closed_or_dense_ker [has_continuous_add M'] [is_simple_module R' R'] - (l : M' →ₗ[R'] R') : - is_closed (l.ker : set M') ∨ dense (l.ker : set M') := -begin - rcases l.surjective_or_eq_zero with (hl|rfl), - { refine l.ker.is_closed_or_dense_of_is_coatom (linear_map.is_coatom_ker_of_surjective hl) }, - { rw linear_map.ker_zero, - left, - exact is_closed_univ }, -end - end closure section pi diff --git a/src/topology/algebra/module/finite_dimension.lean b/src/topology/algebra/module/finite_dimension.lean index 312648ad2a9ca..3a7b1d21c7da2 100644 --- a/src/topology/algebra/module/finite_dimension.lean +++ b/src/topology/algebra/module/finite_dimension.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel, Anatole Dedecker -/ import analysis.locally_convex.balanced_core_hull +import topology.algebra.module.simple import topology.algebra.module.determinant /-! diff --git a/src/topology/algebra/module/simple.lean b/src/topology/algebra/module/simple.lean new file mode 100644 index 0000000000000..befba1c2b38dc --- /dev/null +++ b/src/topology/algebra/module/simple.lean @@ -0,0 +1,35 @@ +/- +Copyright (c) 2022 Anatole Dedecker. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Anatole Dedecker +-/ +import ring_theory.simple_module +import topology.algebra.module.basic + +/-! +# The kernel of a linear function is closed or dense + +In this file we prove (`linear_map.is_closed_or_dense_ker`) that the kernel of a linear function `f +: M →ₗ[R] N` is either closed or dense in `M` provided that `N` is a simple module over `R`. This +applies, e.g., to the case when `R = N` is a division ring. +-/ + +universes u v w + +variables {R : Type u} {M : Type v} {N : Type w} + [ring R] [topological_space R] + [topological_space M] [add_comm_group M] [add_comm_group N] + [module R M] [has_continuous_smul R M] [module R N] + [has_continuous_add M] [is_simple_module R N] + +/-- The kernel of a linear map taking values in a simple module over the base ring is closed or +dense. Applies, e.g., to the case when `R = N` is a division ring. -/ +lemma linear_map.is_closed_or_dense_ker (l : M →ₗ[R] N) : + is_closed (l.ker : set M) ∨ dense (l.ker : set M) := +begin + rcases l.surjective_or_eq_zero with (hl|rfl), + { exact l.ker.is_closed_or_dense_of_is_coatom (linear_map.is_coatom_ker_of_surjective hl) }, + { rw linear_map.ker_zero, + left, + exact is_closed_univ }, +end From 1f4705ccdfe1e557fc54a0ce081a05e33d2e6240 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 18 Mar 2023 18:29:34 +0000 Subject: [PATCH 0657/1291] perf(group_theory/torsion): Speedup `torsion_mul_equiv` (#18614) From 17s to 0.1s on my gitpod. --- src/group_theory/torsion.lean | 10 ++++++++-- 1 file changed, 8 insertions(+), 2 deletions(-) diff --git a/src/group_theory/torsion.lean b/src/group_theory/torsion.lean index 643e4da8e2048..964c95393eead 100644 --- a/src/group_theory/torsion.lean +++ b/src/group_theory/torsion.lean @@ -239,9 +239,15 @@ variable {G} lemma torsion_eq_top (tG : is_torsion G) : torsion G = ⊤ := by ext; tauto /-- A torsion monoid is isomorphic to its torsion submonoid. -/ -@[to_additive "An additive torsion monoid is isomorphic to its torsion submonoid.", simps] +@[to_additive "An additive torsion monoid is isomorphic to its torsion submonoid."] def torsion_mul_equiv (tG : is_torsion G) : torsion G ≃* G := - (mul_equiv.submonoid_congr tG.torsion_eq_top).trans submonoid.top_equiv +(mul_equiv.submonoid_congr tG.torsion_eq_top).trans submonoid.top_equiv + +@[to_additive] lemma torsion_mul_equiv_apply (tG : is_torsion G) (a : torsion G) : + tG.torsion_mul_equiv a = mul_equiv.submonoid_congr tG.torsion_eq_top a := rfl + +@[to_additive] lemma torsion_mul_equiv_symm_apply_coe (tG : is_torsion G) (a : G) : + tG.torsion_mul_equiv.symm a = ⟨submonoid.top_equiv.symm a, tG _⟩ := rfl end monoid.is_torsion From 4e7e7009099d4a88a750de710909b95487bf0124 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sat, 18 Mar 2023 23:36:05 +0000 Subject: [PATCH 0658/1291] chore(linear_algebra/free_module/basic): golf a proof (#18615) We construct the basis in another file, we may as well use it. --- 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 aa89253c1406a..ce8be77ced04b 100644 --- a/src/linear_algebra/free_module/basic.lean +++ b/src/linear_algebra/free_module/basic.lean @@ -8,6 +8,7 @@ import linear_algebra.direct_sum.finsupp import logic.small.basic import linear_algebra.std_basis import linear_algebra.finsupp_vector_space +import linear_algebra.tensor_product_basis /-! @@ -158,8 +159,7 @@ variables [comm_ring R] [add_comm_group M] [module R M] [module.free R M] variables [add_comm_group N] [module R N] [module.free R N] instance tensor : module.free R (M ⊗[R] N) := -of_equiv' (of_equiv' (free.finsupp _ R _) (finsupp_tensor_finsupp' R _ _).symm) - (tensor_product.congr (choose_basis R M).repr (choose_basis R N).repr).symm +let ⟨bM⟩ := exists_basis R M, ⟨bN⟩ := exists_basis R N in of_basis (bM.2.tensor_product bN.2) end comm_ring From 05101c3df9d9cfe9430edc205860c79b6d660102 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sun, 19 Mar 2023 01:36:21 +0000 Subject: [PATCH 0659/1291] feat(algebra/group/commute): `div` lemmas (#18607) `commute` analogs of existing lemmas. Also normalise lemma names about `commute` and `nat.cast`/`int.cast`, following existing `int.cast` lemmas. --- src/algebra/field/basic.lean | 40 ++++++++++++++++++++--------- src/algebra/group/commute.lean | 27 +++++++++++++++++-- src/algebra/group_power/lemmas.lean | 16 +++++------- 3 files changed, 59 insertions(+), 24 deletions(-) diff --git a/src/algebra/field/basic.lean b/src/algebra/field/basic.lean index 0beeca085d3fb..814b4696d87dc 100644 --- a/src/algebra/field/basic.lean +++ b/src/algebra/field/basic.lean @@ -6,7 +6,7 @@ Authors: Robert Lewis, Leonardo de Moura, Johannes Hölzl, Mario Carneiro import algebra.field.defs import algebra.group_with_zero.units.lemmas import algebra.hom.ring -import algebra.ring.inj_surj +import algebra.ring.commute /-! # Lemmas about division (semi)rings and (semi)fields @@ -24,7 +24,7 @@ universe u variables {α β K : Type*} section division_semiring -variables [division_semiring α] {a b c : α} +variables [division_semiring α] {a b c d : α} lemma add_div (a b c : α) : (a + b) / c = a / c + b / c := by simp_rw [div_eq_mul_inv, add_mul] @@ -50,6 +50,18 @@ by rw [add_div, mul_div_cancel _ hc] @[field_simps] lemma div_add' (a b c : α) (hc : c ≠ 0) : a / c + b = (a + b * c) / c := by rwa [add_comm, add_div', add_comm] +protected lemma commute.div_add_div (hbc : commute b c) (hbd : commute b d) (hb : b ≠ 0) + (hd : d ≠ 0) : a / b + c / d = (a * d + b * c) / (b * d) := +by rw [add_div, mul_div_mul_right _ b hd, hbc.eq, hbd.eq, mul_div_mul_right c d hb] + +protected lemma commute.one_div_add_one_div (hab : commute a b) (ha : a ≠ 0) (hb : b ≠ 0) : + 1 / a + 1 / b = (a + b) / (a * b) := +by rw [(commute.one_right a).div_add_div hab ha hb, one_mul, mul_one, add_comm] + +protected lemma commute.inv_add_inv (hab : commute a b) (ha : a ≠ 0) (hb : b ≠ 0) : + a⁻¹ + b⁻¹ = (a + b) / (a * b) := +by rw [inv_eq_one_div, inv_eq_one_div, hab.one_div_add_one_div ha hb] + end division_semiring section division_monoid @@ -94,7 +106,7 @@ by rw neg_inv end division_monoid section division_ring -variables [division_ring K] {a b : K} +variables [division_ring K] {a b c d : K} @[simp] lemma div_neg_self {a : K} (h : a ≠ 0) : a / -a = -1 := by rw [div_neg_eq_neg_div, div_self h] @@ -131,6 +143,14 @@ by rw [(mul_sub_left_distrib (1 / a)), (one_div_mul_cancel ha), mul_sub_right_di instance division_ring.is_domain : is_domain K := no_zero_divisors.to_is_domain _ +protected lemma commute.div_sub_div (hbc : commute b c) (hbd : commute b d) (hb : b ≠ 0) + (hd : d ≠ 0) : a / b - c / d = (a * d - b * c) / (b * d) := +by simpa only [mul_neg, neg_div, ←sub_eq_add_neg] using hbc.neg_right.div_add_div hbd hb hd + +protected lemma commute.inv_sub_inv (hab : commute a b) (ha : a ≠ 0) (hb : b ≠ 0) : + a⁻¹ - b⁻¹ = (b - a) / (a * b) := +by simp only [inv_eq_one_div, (commute.one_right a).div_sub_div hab ha hb, one_mul, mul_one] + end division_ring section semifield @@ -138,13 +158,13 @@ variables [semifield α] {a b c d : α} lemma div_add_div (a : α) (c : α) (hb : b ≠ 0) (hd : d ≠ 0) : (a / b) + (c / d) = ((a * d) + (b * c)) / (b * d) := -by rw [← mul_div_mul_right _ b hd, ← mul_div_mul_left c d hb, div_add_div_same] +(commute.all b _).div_add_div (commute.all _ _) hb hd lemma one_div_add_one_div (ha : a ≠ 0) (hb : b ≠ 0) : 1 / a + 1 / b = (a + b) / (a * b) := -by rw [div_add_div _ _ ha hb, one_mul, mul_one, add_comm] +(commute.all a _).one_div_add_one_div ha hb lemma inv_add_inv (ha : a ≠ 0) (hb : b ≠ 0) : a⁻¹ + b⁻¹ = (a + b) / (a * b) := -by rw [inv_eq_one_div, inv_eq_one_div, one_div_add_one_div ha hb] +(commute.all a _).inv_add_inv ha hb end semifield @@ -156,14 +176,10 @@ local attribute [simp] mul_assoc mul_comm mul_left_comm @[field_simps] lemma div_sub_div (a : K) {b : K} (c : K) {d : K} (hb : b ≠ 0) (hd : d ≠ 0) : (a / b) - (c / d) = ((a * d) - (b * c)) / (b * d) := -begin - simp [sub_eq_add_neg], - rw [neg_eq_neg_one_mul, ← mul_div_assoc, div_add_div _ _ hb hd, - ← mul_assoc, mul_comm b, mul_assoc, ← neg_eq_neg_one_mul] -end +(commute.all b _).div_sub_div (commute.all _ _) hb hd lemma inv_sub_inv {a b : K} (ha : a ≠ 0) (hb : b ≠ 0) : a⁻¹ - b⁻¹ = (b - a) / (a * b) := -by rw [inv_eq_one_div, inv_eq_one_div, div_sub_div _ _ ha hb, one_mul, mul_one] +(commute.all a _).inv_sub_inv ha hb @[field_simps] lemma sub_div' (a b c : K) (hc : c ≠ 0) : b - a / c = (b * c - a) / c := by simpa using div_sub_div b a one_ne_zero hc diff --git a/src/algebra/group/commute.lean b/src/algebra/group/commute.lean index ede1df9536139..b8227479be7cc 100644 --- a/src/algebra/group/commute.lean +++ b/src/algebra/group/commute.lean @@ -85,6 +85,10 @@ by simp only [mul_assoc, h.eq] a * (b * c) = b * (a * c) := by simp only [← mul_assoc, h.eq] +@[to_additive] protected lemma mul_mul_mul_comm (hbc : commute b c) (a d : S) : + (a * b) * (c * d) = (a * c) * (b * d) := +by simp only [hbc.left_comm, mul_assoc] + end semigroup @[to_additive] @@ -171,12 +175,31 @@ u.left_of_mul b a (hc.eq ▸ hu) hc.symm end monoid section division_monoid -variables [division_monoid G] {a b : G} +variables [division_monoid G] {a b c d : G} -@[to_additive] lemma inv_inv : commute a b → commute a⁻¹ b⁻¹ := semiconj_by.inv_inv_symm +@[to_additive] protected lemma inv_inv : commute a b → commute a⁻¹ b⁻¹ := semiconj_by.inv_inv_symm @[simp, to_additive] lemma inv_inv_iff : commute a⁻¹ b⁻¹ ↔ commute a b := semiconj_by.inv_inv_symm_iff +@[to_additive] protected lemma mul_inv (hab : commute a b) : (a * b)⁻¹ = a⁻¹ * b⁻¹ := +by rw [hab.eq, mul_inv_rev] + +@[to_additive] protected lemma inv (hab : commute a b) : (a * b)⁻¹ = a⁻¹ * b⁻¹ := +by rw [hab.eq, mul_inv_rev] + +@[to_additive] protected lemma div_mul_div_comm (hbd : commute b d) (hbc : commute b⁻¹ c) : + a / b * (c / d) = a * c / (b * d) := +by simp_rw [div_eq_mul_inv, mul_inv_rev, hbd.inv_inv.symm.eq, hbc.mul_mul_mul_comm] + +@[to_additive] protected lemma mul_div_mul_comm (hcd : commute c d) (hbc : commute b c⁻¹) : + a * b / (c * d) = a / c * (b / d) := +(hcd.div_mul_div_comm hbc.symm).symm + +@[to_additive] protected lemma div_div_div_comm (hbc : commute b c) (hbd : commute b⁻¹ d) + (hcd : commute c⁻¹ d) : a / b / (c / d) = a / c / (b / d) := +by simp_rw [div_eq_mul_inv, mul_inv_rev, inv_inv, hbd.symm.eq, hcd.symm.eq, + hbc.inv_inv.mul_mul_mul_comm] + end division_monoid section group diff --git a/src/algebra/group_power/lemmas.lean b/src/algebra/group_power/lemmas.lean index 28338034995b0..89259ef81f0ac 100644 --- a/src/algebra/group_power/lemmas.lean +++ b/src/algebra/group_power/lemmas.lean @@ -758,13 +758,12 @@ h.cast_nat_mul_left n commute (m * a : R) (n * b : R) := h.cast_nat_mul_cast_nat_mul m n -@[simp] theorem self_cast_nat_mul (n : ℕ) : commute a (n * a : R) := -(commute.refl a).cast_nat_mul_right n +variables (a) (m n : ℕ) -@[simp] theorem cast_nat_mul_self (n : ℕ) : commute ((n : R) * a) a := -(commute.refl a).cast_nat_mul_left n +@[simp] lemma self_cast_nat_mul : commute a (n * a : R) := (commute.refl a).cast_nat_mul_right n +@[simp] lemma cast_nat_mul_self : commute ((n : R) * a) a := (commute.refl a).cast_nat_mul_left n -@[simp] theorem self_cast_nat_mul_cast_nat_mul (m n : ℕ) : commute (m * a : R) (n * a : R) := +@[simp] theorem self_cast_nat_mul_cast_nat_mul : commute (m * a : R) (n * a : R) := (commute.refl a).cast_nat_mul_cast_nat_mul m n end @@ -792,11 +791,8 @@ h.cast_int_mul_cast_int_mul m n variables (a) (m n : ℤ) -@[simp] lemma cast_int_left : commute (m : R) a := -by { rw [← mul_one (m : R)], exact (one_left a).cast_int_mul_left m } - -@[simp] lemma cast_int_right : commute a m := -by { rw [← mul_one (m : R)], exact (one_right a).cast_int_mul_right m } +@[simp] lemma cast_int_left : commute (m : R) a := int.cast_commute _ _ +@[simp] lemma cast_int_right : commute a m := int.commute_cast _ _ @[simp] theorem self_cast_int_mul : commute a (n * a : R) := (commute.refl a).cast_int_mul_right n From 1dac236edca9b4b6f5f00b1ad831e35f89472837 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Sun, 19 Mar 2023 06:10:09 +0000 Subject: [PATCH 0660/1291] chore(*): add mathlib4 synchronization comments (#18621) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.group_ring_action.invariant` * `algebra.quadratic_discriminant` * `data.dfinsupp.well_founded` * `data.mv_polynomial.comm_ring` * `data.mv_polynomial.invertible` * `data.mv_polynomial.supported` * `data.mv_polynomial.variables` * `data.zmod.basic` * `ring_theory.finiteness` * `ring_theory.multiplicity` * `ring_theory.polynomial.chebyshev` * `ring_theory.valuation.integers` * `topology.category.Born` --- src/algebra/group_ring_action/invariant.lean | 5 ++++- src/algebra/quadratic_discriminant.lean | 3 +++ src/data/dfinsupp/well_founded.lean | 3 +++ src/data/mv_polynomial/comm_ring.lean | 3 +++ src/data/mv_polynomial/invertible.lean | 3 +++ src/data/mv_polynomial/supported.lean | 3 +++ src/data/mv_polynomial/variables.lean | 3 +++ src/data/zmod/basic.lean | 3 +++ src/ring_theory/finiteness.lean | 3 +++ src/ring_theory/multiplicity.lean | 3 +++ src/ring_theory/polynomial/chebyshev.lean | 3 +++ src/ring_theory/valuation/integers.lean | 3 +++ src/topology/category/Born.lean | 3 +++ 13 files changed, 40 insertions(+), 1 deletion(-) diff --git a/src/algebra/group_ring_action/invariant.lean b/src/algebra/group_ring_action/invariant.lean index f0f9a7d16a4a6..83a3c28b1e1e0 100644 --- a/src/algebra/group_ring_action/invariant.lean +++ b/src/algebra/group_ring_action/invariant.lean @@ -6,7 +6,10 @@ Authors: Eric Wieser import algebra.hom.group_action import ring_theory.subring.pointwise -/-! # Subrings invariant under an action -/ +/-! # Subrings invariant under an action + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4.-/ section ring variables (M R : Type*) [monoid M] [ring R] [mul_semiring_action M R] diff --git a/src/algebra/quadratic_discriminant.lean b/src/algebra/quadratic_discriminant.lean index 2821a3d9efe3e..8b13bcb92b846 100644 --- a/src/algebra/quadratic_discriminant.lean +++ b/src/algebra/quadratic_discriminant.lean @@ -12,6 +12,9 @@ import tactic.linear_combination /-! # Quadratic discriminants and roots of a quadratic +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the discriminant of a quadratic and gives the solution to a quadratic equation. ## Main definition diff --git a/src/data/dfinsupp/well_founded.lean b/src/data/dfinsupp/well_founded.lean index f020475406087..a51551fba65f5 100644 --- a/src/data/dfinsupp/well_founded.lean +++ b/src/data/dfinsupp/well_founded.lean @@ -11,6 +11,9 @@ import set_theory.ordinal.basic /-! # Well-foundedness of the lexicographic and product orders on `dfinsupp` and `pi` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The primary results are `dfinsupp.lex.well_founded` and the two variants that follow it, which essentially say that if `(>)` is a well order on `ι`, `(<)` is well-founded on each `α i`, and `0` is a bottom element in `α i`, then the lexicographic `(<)` is well-founded diff --git a/src/data/mv_polynomial/comm_ring.lean b/src/data/mv_polynomial/comm_ring.lean index ab4afa025b694..a7ccd331d6314 100644 --- a/src/data/mv_polynomial/comm_ring.lean +++ b/src/data/mv_polynomial/comm_ring.lean @@ -9,6 +9,9 @@ import data.mv_polynomial.variables /-! # Multivariate polynomials over a ring +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Many results about polynomials hold when the coefficient ring is a commutative semiring. Some stronger results can be derived when we assume this semiring is a ring. diff --git a/src/data/mv_polynomial/invertible.lean b/src/data/mv_polynomial/invertible.lean index cae1d3c1874bf..5d1e6419a1c00 100644 --- a/src/data/mv_polynomial/invertible.lean +++ b/src/data/mv_polynomial/invertible.lean @@ -10,6 +10,9 @@ import ring_theory.algebra_tower /-! # Invertible polynomials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file is a stub containing some basic facts about invertible elements in the ring of polynomials. -/ diff --git a/src/data/mv_polynomial/supported.lean b/src/data/mv_polynomial/supported.lean index 74a8397d4e531..41ce4104cbf5c 100644 --- a/src/data/mv_polynomial/supported.lean +++ b/src/data/mv_polynomial/supported.lean @@ -8,6 +8,9 @@ import data.mv_polynomial.variables /-! # Polynomials supported by a set of variables +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains the definition and lemmas about `mv_polynomial.supported`. ## Main definitions diff --git a/src/data/mv_polynomial/variables.lean b/src/data/mv_polynomial/variables.lean index 7e0447a6f9980..bcbcca71ae509 100644 --- a/src/data/mv_polynomial/variables.lean +++ b/src/data/mv_polynomial/variables.lean @@ -9,6 +9,9 @@ import data.mv_polynomial.rename /-! # Degrees and variables of polynomials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file establishes many results about the degree and variable sets of a multivariate polynomial. The *variable set* of a polynomial $P \in R[X]$ is a `finset` containing each $x \in X$ diff --git a/src/data/zmod/basic.lean b/src/data/zmod/basic.lean index 603f387fe703e..e6472d0723fd0 100644 --- a/src/data/zmod/basic.lean +++ b/src/data/zmod/basic.lean @@ -12,6 +12,9 @@ import tactic.fin_cases /-! # Integers mod `n` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Definition of the integers mod n, and the field structure on the integers mod p. diff --git a/src/ring_theory/finiteness.lean b/src/ring_theory/finiteness.lean index a5ca12ad7a239..bc77fbe165d21 100644 --- a/src/ring_theory/finiteness.lean +++ b/src/ring_theory/finiteness.lean @@ -12,6 +12,9 @@ import ring_theory.ideal.operations /-! # Finiteness conditions in commutative algebra +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define a notion of finiteness that is common in commutative algebra. ## Main declarations diff --git a/src/ring_theory/multiplicity.lean b/src/ring_theory/multiplicity.lean index 3a263db7e015b..f51eca9e5324d 100644 --- a/src/ring_theory/multiplicity.lean +++ b/src/ring_theory/multiplicity.lean @@ -10,6 +10,9 @@ import ring_theory.valuation.basic /-! # Multiplicity of a divisor +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + For a commutative monoid, this file introduces the notion of multiplicity of a divisor and proves several basic results on it. diff --git a/src/ring_theory/polynomial/chebyshev.lean b/src/ring_theory/polynomial/chebyshev.lean index 8c7094a4cc3d3..a62e98c5e4e9a 100644 --- a/src/ring_theory/polynomial/chebyshev.lean +++ b/src/ring_theory/polynomial/chebyshev.lean @@ -9,6 +9,9 @@ import tactic.linear_combination /-! # Chebyshev polynomials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The Chebyshev polynomials are two families of polynomials indexed by `ℕ`, with integral coefficients. diff --git a/src/ring_theory/valuation/integers.lean b/src/ring_theory/valuation/integers.lean index 1f8c31afc0665..f1e917a727c7b 100644 --- a/src/ring_theory/valuation/integers.lean +++ b/src/ring_theory/valuation/integers.lean @@ -9,6 +9,9 @@ import ring_theory.valuation.basic /-! # Ring of integers under a given valuation +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The elements with valuation less than or equal to 1. TODO: Define characteristic predicate. diff --git a/src/topology/category/Born.lean b/src/topology/category/Born.lean index a2b432df926fd..fbc84ec041044 100644 --- a/src/topology/category/Born.lean +++ b/src/topology/category/Born.lean @@ -9,6 +9,9 @@ import topology.bornology.hom /-! # The category of bornologies +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This defines `Born`, the category of bornologies. -/ From c78cad350eb321c81e1eacf68d14e3d3ba1e17f7 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sun, 19 Mar 2023 08:24:29 +0000 Subject: [PATCH 0661/1291] =?UTF-8?q?chore(analysis/inner=5Fproduct=5Fspac?= =?UTF-8?q?e/basic):=20explicit=20`=F0=9D=95=9C`=20argument=20for=20`inner?= =?UTF-8?q?=E2=82=9B=E2=82=97`=20and=20`innerSL`=20(#18613)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit A reasonable fraction of the uses of these functions required either `@` or a type annotation before this change. --- src/analysis/convex/cone/basic.lean | 2 +- src/analysis/inner_product_space/basic.lean | 36 ++++++++++--------- .../inner_product_space/calculus.lean | 2 +- src/analysis/inner_product_space/dual.lean | 6 ++-- .../inner_product_space/l2_space.lean | 2 +- src/analysis/inner_product_space/pi_L2.lean | 2 +- .../inner_product_space/rayleigh.lean | 4 +-- src/analysis/inner_product_space/two_dim.lean | 6 ++-- src/geometry/manifold/instances/sphere.lean | 16 ++++----- src/measure_theory/integral/set_integral.lean | 2 +- 10 files changed, 41 insertions(+), 37 deletions(-) diff --git a/src/analysis/convex/cone/basic.lean b/src/analysis/convex/cone/basic.lean index a619278fbb11a..e31d4ac163d73 100644 --- a/src/analysis/convex/cone/basic.lean +++ b/src/analysis/convex/cone/basic.lean @@ -764,7 +764,7 @@ lemma pointed_inner_dual_cone : s.inner_dual_cone.pointed := /-- The inner dual cone of a singleton is given by the preimage of the positive cone under the linear map `λ y, ⟪x, y⟫`. -/ lemma inner_dual_cone_singleton (x : H) : - ({x} : set H).inner_dual_cone = (convex_cone.positive ℝ ℝ).comap (innerₛₗ x) := + ({x} : set H).inner_dual_cone = (convex_cone.positive ℝ ℝ).comap (innerₛₗ ℝ x) := convex_cone.ext $ λ i, forall_eq lemma inner_dual_cone_union (s t : set H) : diff --git a/src/analysis/inner_product_space/basic.lean b/src/analysis/inner_product_space/basic.lean index 8de1c48ef3df7..1ba001bd04c70 100644 --- a/src/analysis/inner_product_space/basic.lean +++ b/src/analysis/inner_product_space/basic.lean @@ -1695,31 +1695,33 @@ by simp_rw [sum_inner, inner_sum, real_inner_smul_left, real_inner_smul_right, h₁, h₂, zero_mul, mul_zero, finset.sum_const_zero, zero_add, zero_sub, finset.mul_sum, neg_div, finset.sum_div, mul_div_assoc, mul_assoc] +variables (𝕜) + /-- The inner product as a sesquilinear map. -/ def innerₛₗ : E →ₗ⋆[𝕜] E →ₗ[𝕜] 𝕜 := linear_map.mk₂'ₛₗ _ _ (λ v w, ⟪v, w⟫) inner_add_left (λ _ _ _, inner_smul_left _ _ _) inner_add_right (λ _ _ _, inner_smul_right _ _ _) -@[simp] lemma innerₛₗ_apply_coe (v : E) : (innerₛₗ v : E → 𝕜) = λ w, ⟪v, w⟫ := rfl +@[simp] lemma innerₛₗ_apply_coe (v : E) : ⇑(innerₛₗ 𝕜 v) = λ w, ⟪v, w⟫ := rfl -@[simp] lemma innerₛₗ_apply (v w : E) : innerₛₗ v w = ⟪v, w⟫ := rfl +@[simp] lemma innerₛₗ_apply (v w : E) : innerₛₗ 𝕜 v w = ⟪v, w⟫ := rfl /-- The inner product as a continuous sesquilinear map. Note that `to_dual_map` (resp. `to_dual`) in `inner_product_space.dual` is a version of this given as a linear isometry (resp. linear isometric equivalence). -/ def innerSL : E →L⋆[𝕜] E →L[𝕜] 𝕜 := -linear_map.mk_continuous₂ innerₛₗ 1 +linear_map.mk_continuous₂ (innerₛₗ 𝕜) 1 (λ x y, by simp only [norm_inner_le_norm, one_mul, innerₛₗ_apply]) -@[simp] lemma innerSL_apply_coe (v : E) : (innerSL v : E → 𝕜) = λ w, ⟪v, w⟫ := rfl +@[simp] lemma innerSL_apply_coe (v : E) : ⇑(innerSL 𝕜 v) = λ w, ⟪v, w⟫ := rfl -@[simp] lemma innerSL_apply (v w : E) : innerSL v w = ⟪v, w⟫ := rfl +@[simp] lemma innerSL_apply (v w : E) : innerSL 𝕜 v w = ⟪v, w⟫ := rfl /-- `innerSL` is an isometry. Note that the associated `linear_isometry` is defined in `inner_product_space.dual` as `to_dual_map`. -/ -@[simp] lemma innerSL_apply_norm {x : E} : ‖(innerSL x : E →L[𝕜] 𝕜)‖ = ‖x‖ := +@[simp] lemma innerSL_apply_norm (x : E) : ‖innerSL 𝕜 x‖ = ‖x‖ := begin - refine le_antisymm ((innerSL x : E →L[𝕜] 𝕜).op_norm_le_bound (norm_nonneg _) + refine le_antisymm ((innerSL 𝕜 x).op_norm_le_bound (norm_nonneg _) (λ y, norm_inner_le_norm _ _)) _, cases eq_or_lt_of_le (norm_nonneg x) with h h, { have : x = 0 := norm_eq_zero.mp (eq.symm h), @@ -1728,16 +1730,18 @@ begin calc ‖x‖ * ‖x‖ = ‖x‖ ^ 2 : by ring ... = re ⟪x, x⟫ : norm_sq_eq_inner _ ... ≤ abs ⟪x, x⟫ : re_le_abs _ - ... = ‖innerSL x x‖ : by { rw [←is_R_or_C.norm_eq_abs], refl } - ... ≤ ‖innerSL x‖ * ‖x‖ : (innerSL x : E →L[𝕜] 𝕜).le_op_norm _ } + ... = ‖⟪x, x⟫‖ : by rw [←is_R_or_C.norm_eq_abs] + ... ≤ ‖innerSL 𝕜 x‖ * ‖x‖ : (innerSL 𝕜 x).le_op_norm _ } end /-- The inner product as a continuous sesquilinear map, with the two arguments flipped. -/ def innerSL_flip : E →L[𝕜] E →L⋆[𝕜] 𝕜 := @continuous_linear_map.flipₗᵢ' 𝕜 𝕜 𝕜 E E 𝕜 _ _ _ _ _ _ _ _ _ (ring_hom.id 𝕜) (star_ring_end 𝕜) _ _ - innerSL + (innerSL 𝕜) + +@[simp] lemma innerSL_flip_apply (x y : E) : innerSL_flip 𝕜 x y = ⟪y, x⟫ := rfl -@[simp] lemma innerSL_flip_apply (x y : E) : innerSL_flip x y = ⟪y, x⟫ := rfl +variables {𝕜} namespace continuous_linear_map @@ -1748,10 +1752,10 @@ as a continuous linear map. -/ def to_sesq_form : (E →L[𝕜] E') →L[𝕜] E' →L⋆[𝕜] E →L[𝕜] 𝕜 := ↑((continuous_linear_map.flipₗᵢ' E E' 𝕜 (star_ring_end 𝕜) (ring_hom.id 𝕜)).to_continuous_linear_equiv) ∘L -(continuous_linear_map.compSL E E' (E' →L⋆[𝕜] 𝕜) (ring_hom.id 𝕜) (ring_hom.id 𝕜) innerSL_flip) +(continuous_linear_map.compSL E E' (E' →L⋆[𝕜] 𝕜) (ring_hom.id 𝕜) (ring_hom.id 𝕜) (innerSL_flip 𝕜)) @[simp] lemma to_sesq_form_apply_coe (f : E →L[𝕜] E') (x : E') : - to_sesq_form f x = (innerSL x).comp f := rfl + to_sesq_form f x = (innerSL 𝕜 x).comp f := rfl lemma to_sesq_form_apply_norm_le {f : E →L[𝕜] E'} {v : E'} : ‖to_sesq_form f v‖ ≤ ‖f‖ * ‖v‖ := begin @@ -2287,7 +2291,7 @@ by simp [disjoint_iff, K.inf_orthogonal_eq_bot] /-- `Kᗮ` can be characterized as the intersection of the kernels of the operations of inner product with each of the elements of `K`. -/ -lemma orthogonal_eq_inter : Kᗮ = ⨅ v : K, linear_map.ker (innerSL (v:E) : E →L[𝕜] 𝕜) := +lemma orthogonal_eq_inter : Kᗮ = ⨅ v : K, linear_map.ker (innerSL 𝕜 (v : E)) := begin apply le_antisymm, { rw le_infi_iff, @@ -2302,7 +2306,7 @@ end lemma submodule.is_closed_orthogonal : is_closed (Kᗮ : set E) := begin rw orthogonal_eq_inter K, - have := λ v : K, continuous_linear_map.is_closed_ker (innerSL (v:E) : E →L[𝕜] 𝕜), + have := λ v : K, continuous_linear_map.is_closed_ker (innerSL 𝕜 (v : E)), convert is_closed_Inter this, simp only [submodule.infi_coe], end @@ -2397,7 +2401,7 @@ protected lemma continuous_inner : continuous (uncurry inner : completion E × completion E → 𝕜) := begin let inner' : E →+ E →+ 𝕜 := - { to_fun := λ x, (innerₛₗ x).to_add_monoid_hom, + { to_fun := λ x, (innerₛₗ 𝕜 x).to_add_monoid_hom, map_zero' := by ext x; exact inner_zero_left _, map_add' := λ x y, by ext z; exact inner_add_left _ _ _ }, have : continuous (λ p : E × E, inner' p.1 p.2) := continuous_inner, diff --git a/src/analysis/inner_product_space/calculus.lean b/src/analysis/inner_product_space/calculus.lean index 3b2124b70a8be..d146fb60fc3bc 100644 --- a/src/analysis/inner_product_space/calculus.lean +++ b/src/analysis/inner_product_space/calculus.lean @@ -191,7 +191,7 @@ cont_diff_iff_cont_diff_at.2 $ omit 𝕜 lemma has_strict_fderiv_at_norm_sq (x : F) : - has_strict_fderiv_at (λ x, ‖x‖ ^ 2) (bit0 (innerSL x : F →L[ℝ] ℝ)) x := + has_strict_fderiv_at (λ x, ‖x‖ ^ 2) (bit0 (innerSL ℝ x)) x := begin simp only [sq, ← inner_self_eq_norm_mul_norm], convert (has_strict_fderiv_at_id x).inner (has_strict_fderiv_at_id x), diff --git a/src/analysis/inner_product_space/dual.lean b/src/analysis/inner_product_space/dual.lean index 08d37ec9ac3f1..3edc5c7a7c465 100644 --- a/src/analysis/inner_product_space/dual.lean +++ b/src/analysis/inner_product_space/dual.lean @@ -54,14 +54,14 @@ If `E` is complete, this operation is surjective, hence a conjugate-linear isome see `to_dual`. -/ def to_dual_map : E →ₗᵢ⋆[𝕜] normed_space.dual 𝕜 E := -{ norm_map' := λ _, innerSL_apply_norm, - ..innerSL } +{ norm_map' := innerSL_apply_norm _, + ..innerSL 𝕜 } variables {E} @[simp] lemma to_dual_map_apply {x y : E} : to_dual_map 𝕜 E x y = ⟪x, y⟫ := rfl -lemma innerSL_norm [nontrivial E] : ‖(innerSL : E →L⋆[𝕜] E →L[𝕜] 𝕜)‖ = 1 := +lemma innerSL_norm [nontrivial E] : ‖(innerSL 𝕜 : E →L⋆[𝕜] E →L[𝕜] 𝕜)‖ = 1 := show ‖(to_dual_map 𝕜 E).to_continuous_linear_map‖ = 1, from linear_isometry.norm_to_continuous_linear_map _ diff --git a/src/analysis/inner_product_space/l2_space.lean b/src/analysis/inner_product_space/l2_space.lean index 0b5fbdcc5b84c..825030c977a92 100644 --- a/src/analysis/inner_product_space/l2_space.lean +++ b/src/analysis/inner_product_space/l2_space.lean @@ -454,7 +454,7 @@ end protected lemma has_sum_inner_mul_inner (b : hilbert_basis ι 𝕜 E) (x y : E) : has_sum (λ i, ⟪x, b i⟫ * ⟪b i, y⟫) ⟪x, y⟫ := begin - convert (b.has_sum_repr y).mapL (innerSL x), + convert (b.has_sum_repr y).mapL (innerSL _ x), ext i, rw [innerSL_apply, b.repr_apply_apply, inner_smul_right, mul_comm] end diff --git a/src/analysis/inner_product_space/pi_L2.lean b/src/analysis/inner_product_space/pi_L2.lean index d294699838b12..deec2a2192836 100644 --- a/src/analysis/inner_product_space/pi_L2.lean +++ b/src/analysis/inner_product_space/pi_L2.lean @@ -329,7 +329,7 @@ by { simpa using (b.to_basis.equiv_fun_symm_apply v).symm } protected lemma sum_inner_mul_inner (b : orthonormal_basis ι 𝕜 E) (x y : E) : ∑ i, ⟪x, b i⟫ * ⟪b i, y⟫ = ⟪x, y⟫ := begin - have := congr_arg (@innerSL 𝕜 _ _ _ x) (b.sum_repr y), + have := congr_arg (innerSL 𝕜 x) (b.sum_repr y), rw map_sum at this, convert this, ext i, diff --git a/src/analysis/inner_product_space/rayleigh.lean b/src/analysis/inner_product_space/rayleigh.lean index 95fe87c85c53c..9b36279f603f7 100644 --- a/src/analysis/inner_product_space/rayleigh.lean +++ b/src/analysis/inner_product_space/rayleigh.lean @@ -93,7 +93,7 @@ variables {F : Type*} [inner_product_space ℝ F] lemma _root_.linear_map.is_symmetric.has_strict_fderiv_at_re_apply_inner_self {T : F →L[ℝ] F} (hT : (T : F →ₗ[ℝ] F).is_symmetric) (x₀ : F) : - has_strict_fderiv_at T.re_apply_inner_self (_root_.bit0 (innerSL (T x₀) : F →L[ℝ] ℝ)) x₀ := + has_strict_fderiv_at T.re_apply_inner_self (_root_.bit0 (innerSL ℝ (T x₀))) x₀ := begin convert T.has_strict_fderiv_at.inner (has_strict_fderiv_at_id x₀), ext y, @@ -120,7 +120,7 @@ begin refine ⟨a, b, h₁, _⟩, apply (inner_product_space.to_dual_map ℝ F).injective, simp only [linear_isometry.map_add, linear_isometry.map_smul, linear_isometry.map_zero], - change a • innerSL x₀ + b • innerSL (T x₀) = 0, + change a • innerSL _ x₀ + b • innerSL _ (T x₀) = 0, apply smul_right_injective (F →L[ℝ] ℝ) (two_ne_zero : (2:ℝ) ≠ 0), simpa only [_root_.bit0, add_smul, smul_add, one_smul, add_zero] using h₂ end diff --git a/src/analysis/inner_product_space/two_dim.lean b/src/analysis/inner_product_space/two_dim.lean index 96f4ebac6e9ed..2283a8740c7d4 100644 --- a/src/analysis/inner_product_space/two_dim.lean +++ b/src/analysis/inner_product_space/two_dim.lean @@ -369,7 +369,7 @@ coe_basis_of_linear_independent_of_card_eq_finrank _ _ /-- For vectors `a x y : E`, the identity `⟪a, x⟫ * ⟪a, y⟫ + ω a x * ω a y = ‖a‖ ^ 2 * ⟪x, y⟫`. (See `orientation.inner_mul_inner_add_area_form_mul_area_form` for the "applied" form.)-/ lemma inner_mul_inner_add_area_form_mul_area_form' (a x : E) : - ⟪a, x⟫ • @innerₛₗ ℝ _ _ _ a + ω a x • ω a = ‖a‖ ^ 2 • @innerₛₗ ℝ _ _ _ x := + ⟪a, x⟫ • innerₛₗ ℝ a + ω a x • ω a = ‖a‖ ^ 2 • innerₛₗ ℝ x := begin by_cases ha : a = 0, { simp [ha] }, @@ -399,7 +399,7 @@ by simpa [sq, real_inner_self_eq_norm_sq] using o.inner_mul_inner_add_area_form_ /-- For vectors `a x y : E`, the identity `⟪a, x⟫ * ω a y - ω a x * ⟪a, y⟫ = ‖a‖ ^ 2 * ω x y`. (See `orientation.inner_mul_area_form_sub` for the "applied" form.) -/ lemma inner_mul_area_form_sub' (a x : E) : - ⟪a, x⟫ • ω a - ω a x • @innerₛₗ ℝ _ _ _ a = ‖a‖ ^ 2 • ω x := + ⟪a, x⟫ • ω a - ω a x • innerₛₗ ℝ a = ‖a‖ ^ 2 • ω x := begin by_cases ha : a = 0, { simp [ha] }, @@ -456,7 +456,7 @@ real part is the inner product and its imaginary part is `orientation.area_form` On `ℂ` with the standard orientation, `kahler w z = conj w * z`; see `complex.kahler`. -/ def kahler : E →ₗ[ℝ] E →ₗ[ℝ] ℂ := -(linear_map.llcomp ℝ E ℝ ℂ complex.of_real_clm) ∘ₗ (@innerₛₗ ℝ E _ _) +(linear_map.llcomp ℝ E ℝ ℂ complex.of_real_clm) ∘ₗ innerₛₗ ℝ + (linear_map.llcomp ℝ E ℝ ℂ ((linear_map.lsmul ℝ ℂ).flip complex.I)) ∘ₗ ω lemma kahler_apply_apply (x y : E) : o.kahler x y = ⟪x, y⟫ + ω x y • complex.I := rfl diff --git a/src/geometry/manifold/instances/sphere.lean b/src/geometry/manifold/instances/sphere.lean index 45b43f2d94681..113a65c86380f 100644 --- a/src/geometry/manifold/instances/sphere.lean +++ b/src/geometry/manifold/instances/sphere.lean @@ -74,27 +74,27 @@ the orthogonal complement of an element `v` of `E`. It is smooth away from the a through `v` parallel to the orthogonal complement. It restricts on the sphere to the stereographic projection. -/ def stereo_to_fun [complete_space E] (x : E) : (ℝ ∙ v)ᗮ := -(2 / ((1:ℝ) - innerSL v x)) • orthogonal_projection (ℝ ∙ v)ᗮ x +(2 / ((1:ℝ) - innerSL _ v x)) • orthogonal_projection (ℝ ∙ v)ᗮ x variables {v} @[simp] lemma stereo_to_fun_apply [complete_space E] (x : E) : - stereo_to_fun v x = (2 / ((1:ℝ) - innerSL v x)) • orthogonal_projection (ℝ ∙ v)ᗮ x := + stereo_to_fun v x = (2 / ((1:ℝ) - innerSL _ v x)) • orthogonal_projection (ℝ ∙ v)ᗮ x := rfl lemma cont_diff_on_stereo_to_fun [complete_space E] : - cont_diff_on ℝ ⊤ (stereo_to_fun v) {x : E | innerSL v x ≠ (1:ℝ)} := + cont_diff_on ℝ ⊤ (stereo_to_fun v) {x : E | innerSL _ v x ≠ (1:ℝ)} := begin refine cont_diff_on.smul _ (orthogonal_projection ((ℝ ∙ v)ᗮ)).cont_diff.cont_diff_on, refine cont_diff_const.cont_diff_on.div _ _, - { exact (cont_diff_const.sub (innerSL v : E →L[ℝ] ℝ).cont_diff).cont_diff_on }, + { exact (cont_diff_const.sub (innerSL ℝ v).cont_diff).cont_diff_on }, { intros x h h', exact h (sub_eq_zero.mp h').symm } end lemma continuous_on_stereo_to_fun [complete_space E] : - continuous_on (stereo_to_fun v) {x : E | innerSL v x ≠ (1:ℝ)} := + continuous_on (stereo_to_fun v) {x : E | innerSL _ v x ≠ (1:ℝ)} := (@cont_diff_on_stereo_to_fun E _ v _).continuous_on variables (v) @@ -208,7 +208,7 @@ begin ext, simp only [stereo_to_fun_apply, stereo_inv_fun_apply, smul_add], -- name two frequently-occuring quantities and write down their basic properties - set a : ℝ := innerSL v x, + set a : ℝ := innerSL _ v x, set y := orthogonal_projection (ℝ ∙ v)ᗮ x, have split : ↑x = a • v + ↑y, { convert eq_sum_orthogonal_projection_self_orthogonal_complement (ℝ ∙ v) x, @@ -262,8 +262,8 @@ begin orthogonal_projection_orthogonal_complement_singleton_eq_zero v, have h₂ : orthogonal_projection (ℝ ∙ v)ᗮ w = w := orthogonal_projection_mem_subspace_eq_self w, - have h₃ : innerSL v w = (0:ℝ) := submodule.mem_orthogonal_singleton_iff_inner_right.mp w.2, - have h₄ : innerSL v v = (1:ℝ) := by simp [real_inner_self_eq_norm_mul_norm, hv], + have h₃ : innerSL _ v w = (0:ℝ) := submodule.mem_orthogonal_singleton_iff_inner_right.mp w.2, + have h₄ : innerSL _ v v = (1:ℝ) := by simp [real_inner_self_eq_norm_mul_norm, hv], simp [h₁, h₂, h₃, h₄, continuous_linear_map.map_add, continuous_linear_map.map_smul, mul_smul] }, { simp } diff --git a/src/measure_theory/integral/set_integral.lean b/src/measure_theory/integral/set_integral.lean index 37a87b63bcd4d..3e0114cb2705e 100644 --- a/src/measure_theory/integral/set_integral.lean +++ b/src/measure_theory/integral/set_integral.lean @@ -1132,7 +1132,7 @@ local notation `⟪`x`, `y`⟫` := @inner 𝕜 E' _ x y lemma integral_inner {f : α → E'} (hf : integrable f μ) (c : E') : ∫ x, ⟪c, f x⟫ ∂μ = ⟪c, ∫ x, f x ∂μ⟫ := -((@innerSL 𝕜 E' _ _ c).restrict_scalars ℝ).integral_comp_comm hf +((innerSL 𝕜 c).restrict_scalars ℝ).integral_comp_comm hf lemma integral_eq_zero_of_forall_integral_inner_eq_zero (f : α → E') (hf : integrable f μ) (hf_int : ∀ (c : E'), ∫ x, ⟪c, f x⟫ ∂μ = 0) : From e96bdfbd1e8c98a09ff75f7ac6204d142debc840 Mon Sep 17 00:00:00 2001 From: Moritz Doll Date: Sun, 19 Mar 2023 11:16:55 +0000 Subject: [PATCH 0662/1291] chore(linear_algebra/affine_space/affine_subspace) add set_like instance (#18622) Needed for the port. --- .../affine_space/affine_subspace.lean | 20 +++++++++---------- .../affine_space/pointwise.lean | 2 +- 2 files changed, 10 insertions(+), 12 deletions(-) diff --git a/src/linear_algebra/affine_space/affine_subspace.lean b/src/linear_algebra/affine_space/affine_subspace.lean index bb0b5eddee5f1..6649e5d2a692f 100644 --- a/src/linear_algebra/affine_space/affine_subspace.lean +++ b/src/linear_algebra/affine_space/affine_subspace.lean @@ -178,9 +178,9 @@ variables (k : Type*) {V : Type*} (P : Type*) [ring k] [add_comm_group V] [modul [affine_space V P] include V --- TODO Refactor to use `instance : set_like (affine_subspace k P) P :=` instead -instance : has_coe (affine_subspace k P) (set P) := ⟨carrier⟩ -instance : has_mem P (affine_subspace k P) := ⟨λ p s, p ∈ (s : set P)⟩ +instance : set_like (affine_subspace k P) P := +{ coe := carrier, + coe_injective' := λ p q _, by cases p; cases q; congr' } /-- A point is in an affine subspace coerced to a set if and only if it is in that affine subspace. -/ @@ -354,17 +354,15 @@ begin end /-- Two affine subspaces are equal if they have the same points. -/ -@[ext] lemma coe_injective : function.injective (coe : affine_subspace k P → set P) := -λ s1 s2 h, begin - cases s1, - cases s2, - congr, - exact h -end +lemma coe_injective : function.injective (coe : affine_subspace k P → set P) := +set_like.coe_injective + +@[ext] theorem ext {p q : affine_subspace k P} (h : ∀ x, x ∈ p ↔ x ∈ q) : p = q := +set_like.ext h @[simp] lemma ext_iff (s₁ s₂ : affine_subspace k P) : (s₁ : set P) = s₂ ↔ s₁ = s₂ := -⟨λ h, coe_injective h, by tidy⟩ +set_like.ext'_iff.symm /-- Two affine subspaces with the same direction and nonempty intersection are equal. -/ diff --git a/src/linear_algebra/affine_space/pointwise.lean b/src/linear_algebra/affine_space/pointwise.lean index ac31440282c43..8c861c3bc58d7 100644 --- a/src/linear_algebra/affine_space/pointwise.lean +++ b/src/linear_algebra/affine_space/pointwise.lean @@ -49,7 +49,7 @@ lemma vadd_mem_pointwise_vadd_iff {v : V} {s : affine_subspace k P} {p : P} : vadd_mem_vadd_set_iff lemma pointwise_vadd_bot (v : V) : v +ᵥ (⊥ : affine_subspace k P) = ⊥ := -by { ext, simp, } +by simp [set_like.ext'_iff] lemma pointwise_vadd_direction (v : V) (s : affine_subspace k P) : (v +ᵥ s).direction = s.direction := From 7ec294687917cbc5c73620b4414ae9b5dd9ae1b4 Mon Sep 17 00:00:00 2001 From: Michael Stoll Date: Sun, 19 Mar 2023 19:14:00 +0000 Subject: [PATCH 0663/1291] feat(number_theory/pell): group structure (#18568) This continues the sequence of PRs related to the Pell equation. We define a type for the solutions of `x^2 - d*y^2 = 1` and give it the structure of a commutative mutiplicative group with compatible negation. See [this thread on Zulip](https://leanprover.zulipchat.com/#narrow/stream/217875-Is-there-code-for-X.3F/topic/Proving.20Pell's.20equation.20is.20solvable/near/338175458). --- src/number_theory/pell.lean | 106 ++++++++++++++++++++++++++++ src/number_theory/zsqrtd/basic.lean | 13 ++++ 2 files changed, 119 insertions(+) diff --git a/src/number_theory/pell.lean b/src/number_theory/pell.lean index 5c26762c0ac84..656a1950447e8 100644 --- a/src/number_theory/pell.lean +++ b/src/number_theory/pell.lean @@ -7,6 +7,7 @@ Authors: Michael Geißer, Michael Stoll import tactic.qify import data.zmod.basic import number_theory.diophantine_approximation +import number_theory.zsqrtd.basic /-! # Pell's Equation @@ -137,4 +138,109 @@ end end existence +/-! +### Group structure of the solution set + +We define a structure of a commutative multiplicative group with distributive negation +on the set of all solutions to the Pell equation `x^2 - d*y^2 = 1`. + +The type of such solutions is `pell.solution₁ d`. It corresponds to a pair of integers `x` and `y` +and a proof that `(x, y)` is indeed a solution. + +The multiplication is given by `(x, y) * (x', y') = (x*y' + d*y*y', x*y' + y*x')`. +This is obtained by mapping `(x, y)` to `x + y*√d` and multiplying the results. +In fact, we define `pell.solution₁ d` to be `↥(unitary (ℤ√d))` and transport +the "commutative group with distributive negation" structure from `↥(unitary (ℤ√d))`. + +We then set up an API for `pell.solution₁ d`. +-/ + +open zsqrtd + +/-- An element of `ℤ√d` has norm one (i.e., `a.re^2 - d*a.im^2 = 1`) if and only if +it is contained in the submonoid of unitary elements. + +TODO: merge this result with `pell.is_pell_iff_mem_unitary`. -/ +lemma is_pell_solution_iff_mem_unitary {d : ℤ} {a : ℤ√d} : + a.re ^ 2 - d * a.im ^ 2 = 1 ↔ a ∈ unitary ℤ√d := +by rw [← norm_eq_one_iff_mem_unitary, norm_def, sq, sq, ← mul_assoc] + +-- We use `solution₁ d` to allow for a more general structure `solution d m` that +-- encodes solutions to `x^2 - d*y^2 = m` to be added later. + +/-- `pell.solution₁ d` is the type of solutions to the Pell equation `x^2 - d*y^2 = 1`. +We define this in terms of elements of `ℤ√d` of norm one. +-/ +@[derive [comm_group, has_distrib_neg, inhabited]] +def solution₁ (d : ℤ) : Type := ↥(unitary ℤ√d) + +namespace solution₁ + +variables {d : ℤ} + +instance : has_coe (solution₁ d) ℤ√d := { coe := subtype.val } + +/-- The `x` component of a solution to the Pell equation `x^2 - d*y^2 = 1` -/ +protected def x (a : solution₁ d) : ℤ := (a : ℤ√d).re + +/-- The `y` component of a solution to the Pell equation `x^2 - d*y^2 = 1` -/ +protected def y (a : solution₁ d) : ℤ := (a : ℤ√d).im + +/-- The proof that `a` is a solution to the Pell equation `x^2 - d*y^2 = 1` -/ +lemma prop (a : solution₁ d) : a.x ^ 2 - d * a.y ^ 2 = 1 := +is_pell_solution_iff_mem_unitary.mpr a.property + +/-- An alternative form of the equation, suitable for rewriting `x^2`. -/ +lemma prop_x (a : solution₁ d) : a.x ^ 2 = 1 + d * a.y ^ 2 := by {rw ← a.prop, ring} + +/-- An alternative form of the equation, suitable for rewriting `d * y^2`. -/ +lemma prop_y (a : solution₁ d) : d * a.y ^ 2 = a.x ^ 2 - 1 := by {rw ← a.prop, ring} + +/-- Two solutions are equal if their `x` and `y` components are equal. -/ +@[ext] +lemma ext {a b : solution₁ d} (hx : a.x = b.x) (hy : a.y = b.y) : a = b := +subtype.ext $ ext.mpr ⟨hx, hy⟩ + +/-- Construct a solution from `x`, `y` and a proof that the equation is satisfied. -/ +def mk (x y : ℤ) (prop : x ^ 2 - d * y ^ 2 = 1) : solution₁ d := +{ val := ⟨x, y⟩, + property := is_pell_solution_iff_mem_unitary.mp prop } + +@[simp] +lemma x_mk (x y : ℤ) (prop : x ^ 2 - d * y ^ 2 = 1) : (mk x y prop).x = x := rfl + +@[simp] +lemma y_mk (x y : ℤ) (prop : x ^ 2 - d * y ^ 2 = 1) : (mk x y prop).y = y := rfl + +@[simp] +lemma coe_mk (x y : ℤ) (prop : x ^ 2 - d * y ^ 2 = 1) : (↑(mk x y prop) : ℤ√d) = ⟨x,y⟩ := +zsqrtd.ext.mpr ⟨x_mk x y prop, y_mk x y prop⟩ + +@[simp] +lemma x_one : (1 : solution₁ d).x = 1 := rfl + +@[simp] +lemma y_one : (1 : solution₁ d).y = 0 := rfl + +@[simp] +lemma x_mul (a b : solution₁ d) : (a * b).x = a.x * b.x + d * (a.y * b.y) := +by {rw ← mul_assoc, refl} + +@[simp] +lemma y_mul (a b : solution₁ d) : (a * b).y = a.x * b.y + a.y * b.x := rfl + +@[simp] +lemma x_inv (a : solution₁ d) : a⁻¹.x = a.x := rfl + +@[simp] +lemma y_inv (a : solution₁ d) : a⁻¹.y = -a.y := rfl + +@[simp] +lemma x_neg (a : solution₁ d) : (-a).x = -a.x := rfl + +@[simp] +lemma y_neg (a : solution₁ d) : (-a).y = -a.y := rfl + +end solution₁ + end pell diff --git a/src/number_theory/zsqrtd/basic.lean b/src/number_theory/zsqrtd/basic.lean index 3047f61fef974..ea3dbca7b990f 100644 --- a/src/number_theory/zsqrtd/basic.lean +++ b/src/number_theory/zsqrtd/basic.lean @@ -6,6 +6,7 @@ Authors: Mario Carneiro import algebra.associated import ring_theory.int.basic import tactic.ring +import algebra.star.unitary /-! # ℤ[√d] @@ -759,4 +760,16 @@ begin rw [norm_eq_mul_conj, ring_hom.map_mul, ha, zero_mul] end +/-- An element of `ℤ√d` has norm equal to `1` if and only if it is contained in the submonoid +of unitary elements. -/ +lemma norm_eq_one_iff_mem_unitary {d : ℤ} {a : ℤ√d} : a.norm = 1 ↔ a ∈ unitary ℤ√d := +begin + rw [unitary.mem_iff_self_mul_star, ← norm_eq_mul_conj], + norm_cast, +end + +/-- The kernel of the norm map on `ℤ√d` equals the submonoid of unitary elements. -/ +lemma mker_norm_eq_unitary {d : ℤ} : (@norm_monoid_hom d).mker = unitary ℤ√d := +submonoid.ext (λ x, norm_eq_one_iff_mem_unitary) + end zsqrtd From 9abfa6f0727d5adc99067e325e15d1a9de17fd8e Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sun, 19 Mar 2023 20:21:11 +0000 Subject: [PATCH 0664/1291] refactor(linear_algebra/matrix/hermitian): golf and generalize (#18565) A handful of these results can be proven trivially using results about `is_self_adjoint`. This also generalizes the typeclass arguments throughout the file, though largely in a mathematically meaningless way. --- src/algebra/star/pi.lean | 3 + src/algebra/star/prod.lean | 4 + src/algebra/star/self_adjoint.lean | 11 ++ src/linear_algebra/matrix/hermitian.lean | 158 ++++++++++++++--------- 4 files changed, 115 insertions(+), 61 deletions(-) diff --git a/src/algebra/star/pi.lean b/src/algebra/star/pi.lean index 3aafeb402df34..201b46706e46e 100644 --- a/src/algebra/star/pi.lean +++ b/src/algebra/star/pi.lean @@ -29,6 +29,9 @@ instance [Π i, has_star (f i)] : has_star (Π i, f i) := lemma star_def [Π i, has_star (f i)] (x : Π i, f i) : star x = λ i, star (x i) := rfl +instance [Π i, has_star (f i)] [∀ i, has_trivial_star (f i)] : has_trivial_star (Π i, f i) := +{ star_trivial := λ _, funext $ λ _, star_trivial _ } + instance [Π i, has_involutive_star (f i)] : has_involutive_star (Π i, f i) := { star_involutive := λ _, funext $ λ _, star_star _ } diff --git a/src/algebra/star/prod.lean b/src/algebra/star/prod.lean index 3ced92aefa10a..2bc74d1b4e8c8 100644 --- a/src/algebra/star/prod.lean +++ b/src/algebra/star/prod.lean @@ -29,6 +29,10 @@ instance [has_star R] [has_star S] : has_star (R × S) := lemma star_def [has_star R] [has_star S] (x : R × S) : star x = (star x.1, star x.2) := rfl +instance [has_star R] [has_star S] [has_trivial_star R] [has_trivial_star S] : + has_trivial_star (R × S) := +{ star_trivial := λ _, prod.ext (star_trivial _) (star_trivial _) } + instance [has_involutive_star R] [has_involutive_star S] : has_involutive_star (R × S) := { star_involutive := λ _, prod.ext (star_star _) (star_star _) } diff --git a/src/algebra/star/self_adjoint.lean b/src/algebra/star/self_adjoint.lean index 47d431e7355dc..8ebc3e5f2e3af 100644 --- a/src/algebra/star/self_adjoint.lean +++ b/src/algebra/star/self_adjoint.lean @@ -111,6 +111,17 @@ by simp only [is_self_adjoint_iff, star_sub, hx.star_eq, hy.star_eq] end add_group +section add_comm_monoid +variables [add_comm_monoid R] [star_add_monoid R] + +lemma _root_.is_self_adjoint_add_star_self (x : R) : is_self_adjoint (x + star x) := +by simp only [is_self_adjoint_iff, add_comm, star_add, star_star] + +lemma _root_.is_self_adjoint_star_add_self (x : R) : is_self_adjoint (star x + x) := +by simp only [is_self_adjoint_iff, add_comm, star_add, star_star] + +end add_comm_monoid + section semigroup variables [semigroup R] [star_semigroup R] diff --git a/src/linear_algebra/matrix/hermitian.lean b/src/linear_algebra/matrix/hermitian.lean index 4dace90d7d553..9bb0f8bdc780f 100644 --- a/src/linear_algebra/matrix/hermitian.lean +++ b/src/linear_algebra/matrix/hermitian.lean @@ -9,6 +9,8 @@ import analysis.inner_product_space.pi_L2 This file defines hermitian matrices and some basic results about them. +See also `is_self_adjoint`, which generalizes this definition to other star rings. + ## Main definition * `matrix.is_hermitian` : a matrix `A : matrix n n α` is hermitian if `Aᴴ = A`. @@ -27,9 +29,8 @@ open_locale matrix local notation `⟪`x`, `y`⟫` := @inner α _ _ x y -section non_unital_semiring - -variables [non_unital_semiring α] [star_ring α] [non_unital_semiring β] [star_ring β] +section has_star +variables [has_star α] [has_star β] /-- A matrix is hermitian if it is equal to its conjugate transpose. On the reals, this definition captures symmetric matrices. -/ @@ -37,46 +38,19 @@ def is_hermitian (A : matrix n n α) : Prop := Aᴴ = A lemma is_hermitian.eq {A : matrix n n α} (h : A.is_hermitian) : Aᴴ = A := h +protected lemma is_hermitian.is_self_adjoint {A : matrix n n α} (h : A.is_hermitian) : + is_self_adjoint A := h + @[ext] lemma is_hermitian.ext {A : matrix n n α} : (∀ i j, star (A j i) = A i j) → A.is_hermitian := by { intros h, ext i j, exact h i j } lemma is_hermitian.apply {A : matrix n n α} (h : A.is_hermitian) (i j : n) : star (A j i) = A i j := -by { unfold is_hermitian at h, rw [← h, conj_transpose_apply, star_star, h] } +congr_fun (congr_fun h _) _ lemma is_hermitian.ext_iff {A : matrix n n α} : A.is_hermitian ↔ ∀ i j, star (A j i) = A i j := ⟨is_hermitian.apply, is_hermitian.ext⟩ -lemma is_hermitian_mul_conj_transpose_self [fintype n] (A : matrix n n α) : - (A ⬝ Aᴴ).is_hermitian := -by rw [is_hermitian, conj_transpose_mul, conj_transpose_conj_transpose] - -lemma is_hermitian_transpose_mul_self [fintype n] (A : matrix n n α) : - (Aᴴ ⬝ A).is_hermitian := -by rw [is_hermitian, conj_transpose_mul, conj_transpose_conj_transpose] - -lemma is_hermitian_conj_transpose_mul_mul [fintype m] {A : matrix m m α} (B : matrix m n α) - (hA : A.is_hermitian) : (Bᴴ ⬝ A ⬝ B).is_hermitian := -by simp only [is_hermitian, conj_transpose_mul, conj_transpose_conj_transpose, hA.eq, - matrix.mul_assoc] - -lemma is_hermitian_mul_mul_conj_transpose [fintype m] {A : matrix m m α} (B : matrix n m α) - (hA : A.is_hermitian) : (B ⬝ A ⬝ Bᴴ).is_hermitian := -by simp only [is_hermitian, conj_transpose_mul, conj_transpose_conj_transpose, hA.eq, - matrix.mul_assoc] - -lemma is_hermitian_add_transpose_self (A : matrix n n α) : - (A + Aᴴ).is_hermitian := -by simp [is_hermitian, add_comm] - -lemma is_hermitian_transpose_add_self (A : matrix n n α) : - (Aᴴ + A).is_hermitian := -by simp [is_hermitian, add_comm] - -@[simp] lemma is_hermitian_zero : - (0 : matrix n n α).is_hermitian := -conj_transpose_zero - @[simp] lemma is_hermitian.map {A : matrix n n α} (h : A.is_hermitian) (f : α → β) (hf : function.semiconj f star star) : (A.map f).is_hermitian := @@ -95,15 +69,6 @@ lemma is_hermitian.conj_transpose {A : matrix n n α} (h : A.is_hermitian) : Aᴴ.is_hermitian := h.transpose.map _ $ λ _, rfl -@[simp] lemma is_hermitian_conj_transpose_iff (A : matrix n n α) : - Aᴴ.is_hermitian ↔ A.is_hermitian := -⟨by { intro h, rw [← conj_transpose_conj_transpose A], exact is_hermitian.conj_transpose h }, - is_hermitian.conj_transpose⟩ - -@[simp] lemma is_hermitian.add {A B : matrix n n α} (hA : A.is_hermitian) (hB : B.is_hermitian) : - (A + B).is_hermitian := -(conj_transpose_add _ _).trans (hA.symm ▸ hB.symm ▸ rfl) - @[simp] lemma is_hermitian.submatrix {A : matrix n n α} (h : A.is_hermitian) (f : m → n) : (A.submatrix f f).is_hermitian := (conj_transpose_submatrix _ _ _).trans (h.symm ▸ rfl) @@ -112,10 +77,14 @@ h.transpose.map _ $ λ _, rfl (A.submatrix e e).is_hermitian ↔ A.is_hermitian := ⟨λ h, by simpa using h.submatrix e.symm, λ h, h.submatrix _⟩ -/-- The real diagonal matrix `diagonal v` is hermitian. -/ -@[simp] lemma is_hermitian_diagonal [decidable_eq n] (v : n → ℝ) : - (diagonal v).is_hermitian := -diagonal_conj_transpose _ +end has_star + +section has_involutive_star +variables [has_involutive_star α] + +@[simp] lemma is_hermitian_conj_transpose_iff (A : matrix n n α) : + Aᴴ.is_hermitian ↔ A.is_hermitian := +is_self_adjoint.star_iff /-- A block matrix `A.from_blocks B C D` is hermitian, if `A` and `D` are hermitian and `Bᴴ = C`. -/ @@ -124,7 +93,8 @@ lemma is_hermitian.from_blocks (hA : A.is_hermitian) (hBC : Bᴴ = C) (hD : D.is_hermitian) : (A.from_blocks B C D).is_hermitian := begin - have hCB : Cᴴ = B, {rw ← hBC, simp}, + have hCB : Cᴴ = B, + { rw [← hBC, conj_transpose_conj_transpose] }, unfold matrix.is_hermitian, rw from_blocks_conj_transpose, congr; @@ -139,34 +109,100 @@ lemma is_hermitian_from_blocks_iff congr_arg to_blocks₁₂ h, congr_arg to_blocks₂₂ h⟩, λ ⟨hA, hBC, hCB, hD⟩, is_hermitian.from_blocks hA hBC hD⟩ -end non_unital_semiring +end has_involutive_star -section semiring +section add_monoid +variables [add_monoid α] [star_add_monoid α] [add_monoid β] [star_add_monoid β] -variables [semiring α] [star_ring α] [semiring β] [star_ring β] +/-- A diagonal matrix is hermitian if the entries are self-adjoint -/ +lemma is_hermitian_diagonal_of_self_adjoint [decidable_eq n] + (v : n → α) (h : is_self_adjoint v) : + (diagonal v).is_hermitian := +-- TODO: add a `pi.has_trivial_star` instance and remove the `funext` +(diagonal_conj_transpose v).trans $ congr_arg _ h -@[simp] lemma is_hermitian_one [decidable_eq n] : - (1 : matrix n n α).is_hermitian := -conj_transpose_one +/-- A diagonal matrix is hermitian if the entries have the trivial `star` operation +(such as on the reals). -/ +@[simp] lemma is_hermitian_diagonal [has_trivial_star α] [decidable_eq n] (v : n → α) : + (diagonal v).is_hermitian := +is_hermitian_diagonal_of_self_adjoint _ (is_self_adjoint.all _) -end semiring +@[simp] lemma is_hermitian_zero : + (0 : matrix n n α).is_hermitian := +is_self_adjoint_zero _ + +@[simp] lemma is_hermitian.add {A B : matrix n n α} (hA : A.is_hermitian) (hB : B.is_hermitian) : + (A + B).is_hermitian := +is_self_adjoint.add hA hB + +end add_monoid + +section add_comm_monoid +variables [add_comm_monoid α] [star_add_monoid α] + +lemma is_hermitian_add_transpose_self (A : matrix n n α) : + (A + Aᴴ).is_hermitian := +is_self_adjoint_add_star_self A -section ring +lemma is_hermitian_transpose_add_self (A : matrix n n α) : + (Aᴴ + A).is_hermitian := +is_self_adjoint_star_add_self A + +end add_comm_monoid -variables [ring α] [star_ring α] [ring β] [star_ring β] +section add_group +variables [add_group α] [star_add_monoid α] @[simp] lemma is_hermitian.neg {A : matrix n n α} (h : A.is_hermitian) : (-A).is_hermitian := -(conj_transpose_neg _).trans (congr_arg _ h) +is_self_adjoint.neg h @[simp] lemma is_hermitian.sub {A B : matrix n n α} (hA : A.is_hermitian) (hB : B.is_hermitian) : (A - B).is_hermitian := -(conj_transpose_sub _ _).trans (hA.symm ▸ hB.symm ▸ rfl) +is_self_adjoint.sub hA hB -end ring +end add_group -section comm_ring +section non_unital_semiring +variables [non_unital_semiring α] [star_ring α] [non_unital_semiring β] [star_ring β] + +/-- Note this is more general than `is_self_adjoint.mul_star_self` as `B` can be rectangular. -/ +lemma is_hermitian_mul_conj_transpose_self [fintype n] (A : matrix m n α) : + (A ⬝ Aᴴ).is_hermitian := +by rw [is_hermitian, conj_transpose_mul, conj_transpose_conj_transpose] + +/-- Note this is more general than `is_self_adjoint.star_mul_self` as `B` can be rectangular. -/ +lemma is_hermitian_transpose_mul_self [fintype m] (A : matrix m n α) : + (Aᴴ ⬝ A).is_hermitian := +by rw [is_hermitian, conj_transpose_mul, conj_transpose_conj_transpose] +/-- Note this is more general than `is_self_adjoint.conjugate'` as `B` can be rectangular. -/ +lemma is_hermitian_conj_transpose_mul_mul [fintype m] {A : matrix m m α} (B : matrix m n α) + (hA : A.is_hermitian) : (Bᴴ ⬝ A ⬝ B).is_hermitian := +by simp only [is_hermitian, conj_transpose_mul, conj_transpose_conj_transpose, hA.eq, + matrix.mul_assoc] + +/-- Note this is more general than `is_self_adjoint.conjugate` as `B` can be rectangular. -/ +lemma is_hermitian_mul_mul_conj_transpose [fintype m] {A : matrix m m α} (B : matrix n m α) + (hA : A.is_hermitian) : (B ⬝ A ⬝ Bᴴ).is_hermitian := +by simp only [is_hermitian, conj_transpose_mul, conj_transpose_conj_transpose, hA.eq, + matrix.mul_assoc] + +end non_unital_semiring + +section semiring + +variables [semiring α] [star_ring α] [semiring β] [star_ring β] + +/-- Note this is more general for matrices than `is_self_adjoint_one` as it does not +require `fintype n`, which is necessary for `monoid (matrix n n R)`. -/ +@[simp] lemma is_hermitian_one [decidable_eq n] : + (1 : matrix n n α).is_hermitian := +conj_transpose_one + +end semiring + +section comm_ring variables [comm_ring α] [star_ring α] lemma is_hermitian.inv [fintype m] [decidable_eq m] {A : matrix m m α} From 40cc79df33fb2b67e0dabd815d8e4340592e5bff Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sun, 19 Mar 2023 20:21:12 +0000 Subject: [PATCH 0665/1291] feat(data/matrix/kronecker): lemmas about trace, det, and inv (#18610) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The key results here are: * `det (A ⊗ₖ B) = det A ^ fintype.card n * det B ^ fintype.card m` * `trace (A ⊗ₖ B) = trace A * trace B` * `(A ⊗ₖ B)⁻¹ = (A⁻¹ ⊗ₖ B⁻¹)` and the tensor analogues: * `det (A ⊗ₖₜ[R] B) = (det A ^ fintype.card n) ⊗ₜ[R] (det B ^ fintype.card m)` * `trace (A ⊗ₖₜ[R] B) = trace A ⊗ₜ[R] trace B` This also adds some lemmas linking the kronecker product to block diagonal matrices. Generalized from https://github.com/eric-wieser/lean-matrix-cookbook --- src/data/matrix/kronecker.lean | 158 ++++++++++++++++++++++++++++++++- 1 file changed, 156 insertions(+), 2 deletions(-) diff --git a/src/data/matrix/kronecker.lean b/src/data/matrix/kronecker.lean index 0137a0896b01e..5dc1f8bfeada9 100644 --- a/src/data/matrix/kronecker.lean +++ b/src/data/matrix/kronecker.lean @@ -5,6 +5,9 @@ Authors: Filippo A. E. Nuccio, Eric Wieser -/ import data.matrix.basic +import data.matrix.block +import linear_algebra.matrix.determinant +import linear_algebra.matrix.nonsingular_inverse import linear_algebra.tensor_product import ring_theory.tensor_product @@ -114,6 +117,24 @@ begin simp [diagonal, apply_ite f, ite_and, ite_apply, apply_ite (f (a i₁)), hf₁, hf₂], end +lemma kronecker_map_diagonal_right [has_zero β] [has_zero γ] [decidable_eq n] + (f : α → β → γ) (hf : ∀ a, f a 0 = 0) (A : matrix l m α) (b : n → β): + kronecker_map f A (diagonal b) = block_diagonal (λ i, A.map (λ a, f a (b i))) := +begin + ext ⟨i₁, i₂⟩ ⟨j₁, j₂⟩, + simp [diagonal, block_diagonal, apply_ite (f (A i₁ j₁)), hf], +end + +lemma kronecker_map_diagonal_left [has_zero α] [has_zero γ] [decidable_eq l] + (f : α → β → γ) (hf : ∀ b, f 0 b = 0) (a : l → α) (B : matrix m n β) : + kronecker_map f (diagonal a) B = + matrix.reindex (equiv.prod_comm _ _) (equiv.prod_comm _ _) + (block_diagonal (λ i, B.map (λ b, f (a i) b))) := +begin + ext ⟨i₁, i₂⟩ ⟨j₁, j₂⟩, + simp [diagonal, block_diagonal, apply_ite f, ite_apply, hf], +end + @[simp] lemma kronecker_map_one_one [has_zero α] [has_zero β] [has_zero γ] [has_one α] [has_one β] [has_one γ] [decidable_eq m] [decidable_eq n] (f : α → β → γ) (hf₁ : ∀ b, f 0 b = 0) (hf₂ : ∀ a, f a 0 = 0) (hf₃ : f 1 1 = 1) : @@ -182,6 +203,43 @@ begin simp_rw [f.map_sum, linear_map.sum_apply, linear_map.map_sum, h_comm], end +/-- `trace` distributes over `matrix.kronecker_map_bilinear`. + +This is primarily used with `R = ℕ` to prove `matrix.trace_kronecker`. -/ +lemma trace_kronecker_map_bilinear [comm_semiring R] + [fintype m] [fintype n] [add_comm_monoid α] [add_comm_monoid β] [add_comm_monoid γ] + [module R α] [module R β] [module R γ] + (f : α →ₗ[R] β →ₗ[R] γ) (A : matrix m m α) (B : matrix n n β) : + trace (kronecker_map_bilinear f A B) = f (trace A) (trace B) := +by simp_rw [matrix.trace, matrix.diag, kronecker_map_bilinear_apply_apply, + linear_map.map_sum₂, map_sum, ←finset.univ_product_univ, finset.sum_product, kronecker_map] + +/-- `determinant` of `matrix.kronecker_map_bilinear`. + +This is primarily used with `R = ℕ` to prove `matrix.det_kronecker`. -/ +lemma det_kronecker_map_bilinear [comm_semiring R] + [fintype m] [fintype n] [decidable_eq m] [decidable_eq n] [comm_ring α] + [comm_ring β] [comm_ring γ] + [module R α] [module R β] [module R γ] + (f : α →ₗ[R] β →ₗ[R] γ) (h_comm : ∀ a b a' b', f (a * b) (a' * b') = f a a' * f b b') + (A : matrix m m α) (B : matrix n n β) : + det (kronecker_map_bilinear f A B) = + det (A.map (λ a, f a 1)) ^ fintype.card n * det (B.map (λ b, f 1 b)) ^ fintype.card m := +calc det (kronecker_map_bilinear f A B) + = det (kronecker_map_bilinear f A 1 ⬝ kronecker_map_bilinear f 1 B) + : by rw [←kronecker_map_bilinear_mul_mul f h_comm, matrix.mul_one, matrix.one_mul] +... = det (block_diagonal (λ _, A.map (λ a, f a 1))) + * det (block_diagonal (λ _, B.map (λ b, f 1 b))) + : begin + rw [det_mul, ←diagonal_one, ←diagonal_one, + kronecker_map_bilinear_apply_apply, kronecker_map_diagonal_right _ (λ _, _), + kronecker_map_bilinear_apply_apply, kronecker_map_diagonal_left _ (λ _, _), + det_reindex_self], + { exact linear_map.map_zero₂ _ _ }, + { exact map_zero _ }, + end +... = _ : by simp_rw [det_block_diagonal, finset.prod_const, finset.card_univ] + end kronecker_map /-! ### Specialization to `matrix.kronecker_map (*)` -/ @@ -242,10 +300,29 @@ lemma diagonal_kronecker_diagonal [mul_zero_class α] (diagonal a) ⊗ₖ (diagonal b) = diagonal (λ mn, (a mn.1) * (b mn.2)) := kronecker_map_diagonal_diagonal _ zero_mul mul_zero _ _ +lemma kronecker_diagonal [mul_zero_class α] [decidable_eq n] (A : matrix l m α) (b : n → α): + A ⊗ₖ diagonal b = block_diagonal (λ i, mul_opposite.op (b i) • A) := +kronecker_map_diagonal_right _ mul_zero _ _ + +lemma diagonal_kronecker [mul_zero_class α] [decidable_eq l](a : l → α) (B : matrix m n α) : + diagonal a ⊗ₖ B = + matrix.reindex (equiv.prod_comm _ _) (equiv.prod_comm _ _) (block_diagonal (λ i, a i • B)) := +kronecker_map_diagonal_left _ zero_mul _ _ + @[simp] lemma one_kronecker_one [mul_zero_one_class α] [decidable_eq m] [decidable_eq n] : (1 : matrix m m α) ⊗ₖ (1 : matrix n n α) = 1 := kronecker_map_one_one _ zero_mul mul_zero (one_mul _) +lemma kronecker_one [mul_zero_one_class α] [decidable_eq n] (A : matrix l m α) : + A ⊗ₖ (1 : matrix n n α) = block_diagonal (λ i, A) := +(kronecker_diagonal _ _).trans $ congr_arg _ $ funext $ λ _, matrix.ext $ λ _ _, mul_one _ + +lemma one_kronecker [mul_zero_one_class α] [decidable_eq l] (B : matrix m n α) : + (1 : matrix l l α) ⊗ₖ B = + matrix.reindex (equiv.prod_comm _ _) (equiv.prod_comm _ _) (block_diagonal (λ i, B)) := +(diagonal_kronecker _ _).trans $ + congr_arg _ $ congr_arg _ $ funext $ λ _, matrix.ext $ λ _ _, one_mul _ + lemma mul_kronecker_mul [fintype m] [fintype m'] [comm_semiring α] (A : matrix l m α) (B : matrix m n α) (A' : matrix l' m' α) (B' : matrix m' n' α) : (A ⬝ B) ⊗ₖ (A' ⬝ B') = (A ⊗ₖ A') ⬝ (B ⊗ₖ B') := @@ -256,6 +333,49 @@ kronecker_map_bilinear_mul_mul (algebra.lmul ℕ α).to_linear_map mul_mul_mul_c A ⊗ₖ (B ⊗ₖ C) := kronecker_map_assoc₁ _ _ _ _ A B C mul_assoc +lemma trace_kronecker [fintype m] [fintype n] [semiring α] + (A : matrix m m α) (B : matrix n n α) : + trace (A ⊗ₖ B) = trace A * trace B := +trace_kronecker_map_bilinear (algebra.lmul ℕ α).to_linear_map _ _ + +lemma det_kronecker [fintype m] [fintype n] [decidable_eq m] [decidable_eq n] [comm_ring R] + (A : matrix m m R) (B : matrix n n R) : + det (A ⊗ₖ B) = det A ^ fintype.card n * det B ^ fintype.card m := +begin + refine + (det_kronecker_map_bilinear (algebra.lmul ℕ R).to_linear_map mul_mul_mul_comm _ _).trans _, + congr' 3, + { ext i j, exact mul_one _}, + { ext i j, exact one_mul _}, +end + +lemma inv_kronecker [fintype m] [fintype n] [decidable_eq m] [decidable_eq n] [comm_ring R] + (A : matrix m m R) (B : matrix n n R) : + (A ⊗ₖ B)⁻¹ = A⁻¹ ⊗ₖ B⁻¹ := +begin + -- handle the special cases where either matrix is not invertible + by_cases hA : is_unit A.det, swap, + { casesI is_empty_or_nonempty n, + { exact subsingleton.elim _ _ }, + have hAB : ¬is_unit (A ⊗ₖ B).det, + { refine mt (λ hAB, _) hA, + rw det_kronecker at hAB, + exact (is_unit_pow_iff fintype.card_ne_zero).mp (is_unit_of_mul_is_unit_left hAB) }, + rw [nonsing_inv_apply_not_is_unit _ hA, zero_kronecker, nonsing_inv_apply_not_is_unit _ hAB] }, + by_cases hB : is_unit B.det, swap, + { casesI is_empty_or_nonempty m, + { exact subsingleton.elim _ _ }, + have hAB : ¬is_unit (A ⊗ₖ B).det, + { refine mt (λ hAB, _) hB, + rw det_kronecker at hAB, + exact (is_unit_pow_iff fintype.card_ne_zero).mp (is_unit_of_mul_is_unit_right hAB) }, + rw [nonsing_inv_apply_not_is_unit _ hB, kronecker_zero, + nonsing_inv_apply_not_is_unit _ hAB] }, + -- otherwise follows trivially from `mul_kronecker_mul` + { apply inv_eq_right_inv, + rw [←mul_kronecker_mul, ←one_kronecker_one, mul_nonsing_inv _ hA, mul_nonsing_inv _ hB] }, +end + end kronecker /-! ### Specialization to `matrix.kronecker_map (⊗ₜ)` -/ @@ -324,19 +444,34 @@ lemma diagonal_kronecker_tmul_diagonal (diagonal a) ⊗ₖₜ[R] (diagonal b) = diagonal (λ mn, a mn.1 ⊗ₜ b mn.2) := kronecker_map_diagonal_diagonal _ (zero_tmul _) (tmul_zero _) _ _ +lemma kronecker_tmul_diagonal [decidable_eq n] (A : matrix l m α) (b : n → α): + A ⊗ₖₜ[R] (diagonal b) = block_diagonal (λ i, A.map (λ a, a ⊗ₜ[R] b i)) := +kronecker_map_diagonal_right _ (tmul_zero _) _ _ + +lemma diagonal_kronecker_tmul [decidable_eq l](a : l → α) (B : matrix m n α) : + diagonal a ⊗ₖₜ[R] B = + matrix.reindex (equiv.prod_comm _ _) (equiv.prod_comm _ _) + (block_diagonal (λ i, B.map (λ b, a i ⊗ₜ[R] b))) := +kronecker_map_diagonal_left _ (zero_tmul _) _ _ + @[simp] lemma kronecker_tmul_assoc (A : matrix l m α) (B : matrix n p β) (C : matrix q r γ) : reindex (equiv.prod_assoc l n q) (equiv.prod_assoc m p r) (((A ⊗ₖₜ[R] B) ⊗ₖₜ[R] C).map (tensor_product.assoc _ _ _ _)) = A ⊗ₖₜ[R] (B ⊗ₖₜ[R] C) := ext $ λ i j, assoc_tmul _ _ _ +lemma trace_kronecker_tmul [fintype m] [fintype n] (A : matrix m m α) (B : matrix n n β) : + trace (A ⊗ₖₜ[R] B) = trace A ⊗ₜ[R] trace B := +trace_kronecker_map_bilinear (tensor_product.mk R α β) _ _ + end module section algebra -variables [comm_semiring R] [semiring α] [semiring β] [algebra R α] [algebra R β] - open_locale kronecker open algebra.tensor_product +section semiring +variables [comm_semiring R] [semiring α] [semiring β] [algebra R α] [algebra R β] + @[simp] lemma one_kronecker_tmul_one [decidable_eq m] [decidable_eq n] : (1 : matrix m m α) ⊗ₖₜ[R] (1 : matrix n n α) = 1 := kronecker_map_one_one _ (zero_tmul _) (tmul_zero _) rfl @@ -346,6 +481,25 @@ lemma mul_kronecker_tmul_mul [fintype m] [fintype m'] (A ⬝ B) ⊗ₖₜ[R] (A' ⬝ B') = (A ⊗ₖₜ A') ⬝ (B ⊗ₖₜ B') := kronecker_map_bilinear_mul_mul (tensor_product.mk R α β) tmul_mul_tmul A B A' B' +end semiring + +section comm_ring +variables [comm_ring R] [comm_ring α] [comm_ring β] [algebra R α] [algebra R β] + +lemma det_kronecker_tmul [fintype m] [fintype n] [decidable_eq m] [decidable_eq n] + (A : matrix m m α) (B : matrix n n β) : + det (A ⊗ₖₜ[R] B) = (det A ^ fintype.card n) ⊗ₜ[R] (det B ^ fintype.card m) := +begin + refine + (det_kronecker_map_bilinear (tensor_product.mk R α β) tmul_mul_tmul _ _).trans _, + simp only [mk_apply, ←include_left_apply, ←include_right_apply] {eta := ff}, + simp only [←alg_hom.map_matrix_apply, ←alg_hom.map_det], + simp only [include_left_apply, include_right_apply, tmul_pow, tmul_mul_tmul, + one_pow, _root_.mul_one, _root_.one_mul], +end + +end comm_ring + end algebra -- insert lemmas specific to `kronecker_tmul` below this line From 290a7ba01fbcab1b64757bdaa270d28f4dcede35 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Mon, 20 Mar 2023 06:49:15 +0000 Subject: [PATCH 0666/1291] chore(*): add mathlib4 synchronization comments (#18623) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.dual_number` * `data.finsupp.well_founded` * `data.nat.multiplicity` * `data.polynomial.div` * `data.zmod.algebra` * `data.zmod.parity` * `ring_theory.ideal.idempotent_fg` * `ring_theory.mv_polynomial.symmetric` --- src/algebra/dual_number.lean | 3 +++ src/data/finsupp/well_founded.lean | 3 +++ src/data/nat/multiplicity.lean | 3 +++ src/data/polynomial/div.lean | 3 +++ src/data/zmod/algebra.lean | 3 +++ src/data/zmod/parity.lean | 3 +++ src/ring_theory/ideal/idempotent_fg.lean | 3 +++ src/ring_theory/mv_polynomial/symmetric.lean | 3 +++ 8 files changed, 24 insertions(+) diff --git a/src/algebra/dual_number.lean b/src/algebra/dual_number.lean index 9daad87c0525c..31cacf6c8cbeb 100644 --- a/src/algebra/dual_number.lean +++ b/src/algebra/dual_number.lean @@ -9,6 +9,9 @@ import algebra.triv_sq_zero_ext /-! # Dual numbers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The dual numbers over `R` are of the form `a + bε`, where `a` and `b` are typically elements of a commutative ring `R`, and `ε` is a symbol satisfying `ε^2 = 0`. They are a special case of `triv_sq_zero_ext R M` with `M = R`. diff --git a/src/data/finsupp/well_founded.lean b/src/data/finsupp/well_founded.lean index a5099ed0ffbf7..9a9450b297ebc 100644 --- a/src/data/finsupp/well_founded.lean +++ b/src/data/finsupp/well_founded.lean @@ -9,6 +9,9 @@ import data.finsupp.lex /-! # Well-foundedness of the lexicographic and product orders on `finsupp` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + `finsupp.lex.well_founded` and the two variants that follow it essentially say that if `(>)` is a well order on `α`, `(<)` is well-founded on `N`, and `0` is a bottom element in `N`, then the lexicographic `(<)` is well-founded on `α →₀ N`. diff --git a/src/data/nat/multiplicity.lean b/src/data/nat/multiplicity.lean index 68a56270c1855..f54f1afac1c5b 100644 --- a/src/data/nat/multiplicity.lean +++ b/src/data/nat/multiplicity.lean @@ -14,6 +14,9 @@ import ring_theory.multiplicity /-! # Natural number multiplicity +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains lemmas about the multiplicity function (the maximum prime power dividing a number) when applied to naturals, in particular calculating it for factorials and binomial coefficients. diff --git a/src/data/polynomial/div.lean b/src/data/polynomial/div.lean index 67ab0564bd65b..9abd9e5cf5e15 100644 --- a/src/data/polynomial/div.lean +++ b/src/data/polynomial/div.lean @@ -11,6 +11,9 @@ import ring_theory.multiplicity /-! # Division of univariate polynomials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The main defs are `div_by_monic` and `mod_by_monic`. The compatibility between these is given by `mod_by_monic_add_div`. We also define `root_multiplicity`. diff --git a/src/data/zmod/algebra.lean b/src/data/zmod/algebra.lean index 7646d996b20cd..65f3f96fc884f 100644 --- a/src/data/zmod/algebra.lean +++ b/src/data/zmod/algebra.lean @@ -9,6 +9,9 @@ import algebra.algebra.basic /-! # The `zmod n`-algebra structure on rings whose characteristic divides `n` + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ namespace zmod diff --git a/src/data/zmod/parity.lean b/src/data/zmod/parity.lean index f08e23f86efa9..c83d693c6c394 100644 --- a/src/data/zmod/parity.lean +++ b/src/data/zmod/parity.lean @@ -8,6 +8,9 @@ import data.zmod.basic /-! # Relating parity to natural numbers mod 2 +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This module provides lemmas relating `zmod 2` to `even` and `odd`. ## Tags diff --git a/src/ring_theory/ideal/idempotent_fg.lean b/src/ring_theory/ideal/idempotent_fg.lean index 029aa6c03d62a..61bdcf8506b19 100644 --- a/src/ring_theory/ideal/idempotent_fg.lean +++ b/src/ring_theory/ideal/idempotent_fg.lean @@ -8,6 +8,9 @@ import ring_theory.finiteness /-! ## Lemmas on idempotent finitely generated ideals + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ namespace ideal diff --git a/src/ring_theory/mv_polynomial/symmetric.lean b/src/ring_theory/mv_polynomial/symmetric.lean index d3bf7caf12f6b..4922a8cfafbc2 100644 --- a/src/ring_theory/mv_polynomial/symmetric.lean +++ b/src/ring_theory/mv_polynomial/symmetric.lean @@ -10,6 +10,9 @@ import algebra.algebra.subalgebra.basic /-! # Symmetric Polynomials and Elementary Symmetric Polynomials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines symmetric `mv_polynomial`s and elementary symmetric `mv_polynomial`s. We also prove some basic facts about them. From 4601791ea62fea875b488dafc4e6dede19e8363f Mon Sep 17 00:00:00 2001 From: sgouezel Date: Tue, 21 Mar 2023 06:51:36 +0000 Subject: [PATCH 0667/1291] feat(analysis/calculus/cont_diff): formula for iterated derivative of composition with linear maps (#18592) And missing API for derivatives and iterated derivatives. --- src/analysis/calculus/cont_diff.lean | 234 ++++++++++++++++-- src/analysis/calculus/fderiv.lean | 120 ++++++++- .../normed_space/linear_isometry.lean | 5 + src/analysis/normed_space/multilinear.lean | 116 ++++++++- src/analysis/normed_space/operator_norm.lean | 20 +- src/topology/algebra/module/multilinear.lean | 1 - 6 files changed, 463 insertions(+), 33 deletions(-) diff --git a/src/analysis/calculus/cont_diff.lean b/src/analysis/calculus/cont_diff.lean index 5274779ceabab..37d469b3cc08a 100644 --- a/src/analysis/calculus/cont_diff.lean +++ b/src/analysis/calculus/cont_diff.lean @@ -164,7 +164,7 @@ open_locale classical big_operators nnreal local notation `∞` := (⊤ : ℕ∞) -universes u v w +universes u v w uE uF uG local attribute [instance, priority 1001] normed_add_comm_group.to_add_comm_group normed_space.to_module' add_comm_group.to_add_comm_monoid @@ -173,9 +173,9 @@ open set fin filter function open_locale topology variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] -{E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] -{F : Type*} [normed_add_comm_group F] [normed_space 𝕜 F] -{G : Type*} [normed_add_comm_group G] [normed_space 𝕜 G] +{E : Type uE} [normed_add_comm_group E] [normed_space 𝕜 E] +{F : Type uF} [normed_add_comm_group F] [normed_space 𝕜 F] +{G : Type uG} [normed_add_comm_group G] [normed_space 𝕜 G] {X : Type*} [normed_add_comm_group X] [normed_space 𝕜 X] {s s₁ t u : set E} {f f₁ : E → F} {g : F → G} {x x₀ : E} {c : F} {b : E × F → G} {m n : ℕ∞} @@ -772,7 +772,7 @@ variable {𝕜} lemma iterated_fderiv_within_zero_eq_comp : iterated_fderiv_within 𝕜 0 f s = (continuous_multilinear_curry_fin0 𝕜 E F).symm ∘ f := rfl -lemma norm_iterated_fderiv_within_zero : +@[simp] lemma norm_iterated_fderiv_within_zero : ‖iterated_fderiv_within 𝕜 0 f s x‖ = ‖f x‖ := by rw [iterated_fderiv_within_zero_eq_comp, linear_isometry_equiv.norm_map] @@ -1423,7 +1423,7 @@ variable {𝕜} lemma iterated_fderiv_zero_eq_comp : iterated_fderiv 𝕜 0 f = (continuous_multilinear_curry_fin0 𝕜 E F).symm ∘ f := rfl -lemma norm_iterated_fderiv_zero : +@[simp] lemma norm_iterated_fderiv_zero : ‖iterated_fderiv 𝕜 0 f x‖ = ‖f x‖ := by rw [iterated_fderiv_zero_eq_comp, linear_isometry_equiv.norm_map] @@ -1533,7 +1533,7 @@ by { rw [iterated_fderiv_succ_apply_right, iterated_fderiv_zero_apply], refl } /-- When a function is `C^n` in a set `s` of unique differentiability, it admits `ftaylor_series_within 𝕜 f s` as a Taylor series up to order `n` in `s`. -/ -theorem cont_diff_on_iff_ftaylor_series : +theorem cont_diff_iff_ftaylor_series : cont_diff 𝕜 n f ↔ has_ftaylor_series_up_to n f (ftaylor_series 𝕜 f) := begin split, @@ -1765,6 +1765,99 @@ lemma cont_diff.continuous_linear_map_comp {f : E → F} (g : F →L[𝕜] G) cont_diff_on_univ.1 $ cont_diff_on.continuous_linear_map_comp _ (cont_diff_on_univ.2 hf) +/-- The iterated derivative within a set of the composition with a linear map on the left is +obtained by applying the linear map to the iterated derivative. -/ +lemma continuous_linear_map.iterated_fderiv_within_comp_left + {f : E → F} (g : F →L[𝕜] G) (hf : cont_diff_on 𝕜 n f s) (hs : unique_diff_on 𝕜 s) (hx : x ∈ s) + {i : ℕ} (hi : (i : with_top ℕ) ≤ n) : + iterated_fderiv_within 𝕜 i (g ∘ f) s x = + g.comp_continuous_multilinear_map (iterated_fderiv_within 𝕜 i f s x) := +(((hf.ftaylor_series_within hs).continuous_linear_map_comp g).eq_ftaylor_series_of_unique_diff_on + hi hs hx).symm + +/-- The iterated derivative of the composition with a linear map on the left is +obtained by applying the linear map to the iterated derivative. -/ +lemma continuous_linear_map.iterated_fderiv_comp_left + {f : E → F} (g : F →L[𝕜] G) (hf : cont_diff 𝕜 n f) (x : E) {i : ℕ} (hi : (i : with_top ℕ) ≤ n) : + iterated_fderiv 𝕜 i (g ∘ f) x = g.comp_continuous_multilinear_map (iterated_fderiv 𝕜 i f x) := +begin + simp only [← iterated_fderiv_within_univ], + exact g.iterated_fderiv_within_comp_left hf.cont_diff_on unique_diff_on_univ (mem_univ x) hi, +end + +/-- The iterated derivative within a set of the composition with a linear equiv on the left is +obtained by applying the linear equiv to the iterated derivative. This is true without +differentiability assumptions. -/ +lemma continuous_linear_equiv.iterated_fderiv_within_comp_left + (g : F ≃L[𝕜] G) (f : E → F) (hs : unique_diff_on 𝕜 s) (hx : x ∈ s) (i : ℕ) : + iterated_fderiv_within 𝕜 i (g ∘ f) s x = + (g : F →L[𝕜] G).comp_continuous_multilinear_map (iterated_fderiv_within 𝕜 i f s x) := +begin + induction i with i IH generalizing x, + { ext1 m, + simp only [iterated_fderiv_within_zero_apply, continuous_linear_equiv.coe_coe, + continuous_linear_map.comp_continuous_multilinear_map_coe, embedding_like.apply_eq_iff_eq] }, + { ext1 m, + rw iterated_fderiv_within_succ_apply_left, + have Z : fderiv_within 𝕜 (iterated_fderiv_within 𝕜 i (g ∘ f) s) s x = + fderiv_within 𝕜 (λ y, g.comp_continuous_multilinear_mapL (λ (j : fin i), E) + (iterated_fderiv_within 𝕜 i f s y)) s x, + from fderiv_within_congr' (hs x hx) (λ y hy, IH hy) hx, + simp_rw Z, + rw (g.comp_continuous_multilinear_mapL (λ (j : fin i), E)).comp_fderiv_within (hs x hx), + simp only [continuous_linear_map.coe_comp', continuous_linear_equiv.coe_coe, comp_app, + continuous_linear_equiv.comp_continuous_multilinear_mapL_apply, + continuous_linear_map.comp_continuous_multilinear_map_coe, embedding_like.apply_eq_iff_eq], + rw iterated_fderiv_within_succ_apply_left } +end + +/-- Composition with a linear isometry on the left preserves the norm of the iterated +derivative within a set. -/ +lemma linear_isometry.norm_iterated_fderiv_within_comp_left + {f : E → F} (g : F →ₗᵢ[𝕜] G) (hf : cont_diff_on 𝕜 n f s) (hs : unique_diff_on 𝕜 s) (hx : x ∈ s) + {i : ℕ} (hi : (i : with_top ℕ) ≤ n) : + ‖iterated_fderiv_within 𝕜 i (g ∘ f) s x‖ = ‖iterated_fderiv_within 𝕜 i f s x‖ := +begin + have : iterated_fderiv_within 𝕜 i (g ∘ f) s x = + g.to_continuous_linear_map.comp_continuous_multilinear_map (iterated_fderiv_within 𝕜 i f s x), + from g.to_continuous_linear_map.iterated_fderiv_within_comp_left hf hs hx hi, + rw this, + apply linear_isometry.norm_comp_continuous_multilinear_map +end + +/-- Composition with a linear isometry on the left preserves the norm of the iterated +derivative. -/ +lemma linear_isometry.norm_iterated_fderiv_comp_left + {f : E → F} (g : F →ₗᵢ[𝕜] G) (hf : cont_diff 𝕜 n f) (x : E) {i : ℕ} (hi : (i : with_top ℕ) ≤ n) : + ‖iterated_fderiv 𝕜 i (g ∘ f) x‖ = ‖iterated_fderiv 𝕜 i f x‖ := +begin + simp only [← iterated_fderiv_within_univ], + exact g.norm_iterated_fderiv_within_comp_left hf.cont_diff_on unique_diff_on_univ (mem_univ x) hi +end + +/-- Composition with a linear isometry equiv on the left preserves the norm of the iterated +derivative within a set. -/ +lemma linear_isometry_equiv.norm_iterated_fderiv_within_comp_left + (g : F ≃ₗᵢ[𝕜] G) (f : E → F) (hs : unique_diff_on 𝕜 s) (hx : x ∈ s) (i : ℕ) : + ‖iterated_fderiv_within 𝕜 i (g ∘ f) s x‖ = ‖iterated_fderiv_within 𝕜 i f s x‖ := +begin + have : iterated_fderiv_within 𝕜 i (g ∘ f) s x = + (g : F →L[𝕜] G).comp_continuous_multilinear_map (iterated_fderiv_within 𝕜 i f s x), + from g.to_continuous_linear_equiv.iterated_fderiv_within_comp_left f hs hx i, + rw this, + apply linear_isometry.norm_comp_continuous_multilinear_map g.to_linear_isometry, +end + +/-- Composition with a linear isometry equiv on the left preserves the norm of the iterated +derivative. -/ +lemma linear_isometry_equiv.norm_iterated_fderiv_comp_left + (g : F ≃ₗᵢ[𝕜] G) (f : E → F) (x : E) (i : ℕ) : + ‖iterated_fderiv 𝕜 i (g ∘ f) x‖ = ‖iterated_fderiv 𝕜 i f x‖ := +begin + rw [← iterated_fderiv_within_univ, ← iterated_fderiv_within_univ], + apply g.norm_iterated_fderiv_within_comp_left f unique_diff_on_univ (mem_univ x) i +end + /-- Composition by continuous linear equivs on the left respects higher differentiability at a point in a domain. -/ lemma continuous_linear_equiv.comp_cont_diff_within_at_iff @@ -1849,6 +1942,78 @@ lemma cont_diff.comp_continuous_linear_map {f : E → F} {g : G →L[𝕜] E} cont_diff_on_univ.1 $ cont_diff_on.comp_continuous_linear_map (cont_diff_on_univ.2 hf) _ +/-- The iterated derivative within a set of the composition with a linear map on the right is +obtained by composing the iterated derivative with the linear map. -/ +lemma continuous_linear_map.iterated_fderiv_within_comp_right + {f : E → F} (g : G →L[𝕜] E) (hf : cont_diff_on 𝕜 n f s) + (hs : unique_diff_on 𝕜 s) (h's : unique_diff_on 𝕜 (g⁻¹' s)) {x : G} + (hx : g x ∈ s) {i : ℕ} (hi : (i : with_top ℕ) ≤ n) : + iterated_fderiv_within 𝕜 i (f ∘ g) (g ⁻¹' s) x = + (iterated_fderiv_within 𝕜 i f s (g x)).comp_continuous_linear_map (λ _, g) := +(((hf.ftaylor_series_within hs).comp_continuous_linear_map g).eq_ftaylor_series_of_unique_diff_on + hi h's hx).symm + +/-- The iterated derivative within a set of the composition with a linear equiv on the right is +obtained by composing the iterated derivative with the linear equiv. -/ +lemma continuous_linear_equiv.iterated_fderiv_within_comp_right + (g : G ≃L[𝕜] E) (f : E → F) (hs : unique_diff_on 𝕜 s) {x : G} (hx : g x ∈ s) (i : ℕ) : + iterated_fderiv_within 𝕜 i (f ∘ g) (g ⁻¹' s) x = + (iterated_fderiv_within 𝕜 i f s (g x)).comp_continuous_linear_map (λ _, g) := +begin + induction i with i IH generalizing x, + { ext1 m, + simp only [iterated_fderiv_within_zero_apply, + continuous_multilinear_map.comp_continuous_linear_map_apply] }, + { ext1 m, + simp only [continuous_multilinear_map.comp_continuous_linear_map_apply, + continuous_linear_equiv.coe_coe, iterated_fderiv_within_succ_apply_left], + have : fderiv_within 𝕜 (iterated_fderiv_within 𝕜 i (f ∘ ⇑g) (⇑g ⁻¹' s)) (⇑g ⁻¹' s) x + = fderiv_within 𝕜 (λ y, continuous_multilinear_map.comp_continuous_linear_map_equivL _ + (λ (_x : fin i), g) (iterated_fderiv_within 𝕜 i f s (g y))) (g ⁻¹' s) x, + from fderiv_within_congr' (g.unique_diff_on_preimage_iff.2 hs x hx) (λ y hy, IH hy) hx, + rw [this], + rw continuous_linear_equiv.comp_fderiv_within _ (g.unique_diff_on_preimage_iff.2 hs x hx), + simp only [continuous_linear_map.coe_comp', continuous_linear_equiv.coe_coe, comp_app, + continuous_multilinear_map.comp_continuous_linear_map_equivL_apply, + continuous_multilinear_map.comp_continuous_linear_map_apply], + rw continuous_linear_equiv.comp_right_fderiv_within _ (g.unique_diff_on_preimage_iff.2 hs x hx), + refl } +end + +/-- The iterated derivative of the composition with a linear map on the right is +obtained by composing the iterated derivative with the linear map. -/ +lemma continuous_linear_map.iterated_fderiv_comp_right + (g : G →L[𝕜] E) {f : E → F} (hf : cont_diff 𝕜 n f) (x : G) {i : ℕ} (hi : (i : with_top ℕ) ≤ n) : + iterated_fderiv 𝕜 i (f ∘ g) x = + (iterated_fderiv 𝕜 i f (g x)).comp_continuous_linear_map (λ _, g) := +begin + simp only [← iterated_fderiv_within_univ], + apply g.iterated_fderiv_within_comp_right hf.cont_diff_on unique_diff_on_univ unique_diff_on_univ + (mem_univ _) hi, +end + +/-- Composition with a linear isometry on the right preserves the norm of the iterated derivative +within a set. -/ +lemma linear_isometry_equiv.norm_iterated_fderiv_within_comp_right + (g : G ≃ₗᵢ[𝕜] E) (f : E → F) (hs : unique_diff_on 𝕜 s) {x : G} (hx : g x ∈ s) (i : ℕ) : + ‖iterated_fderiv_within 𝕜 i (f ∘ g) (g ⁻¹' s) x‖ = ‖iterated_fderiv_within 𝕜 i f s (g x)‖ := +begin + have : iterated_fderiv_within 𝕜 i (f ∘ g) (g ⁻¹' s) x = + (iterated_fderiv_within 𝕜 i f s (g x)).comp_continuous_linear_map (λ _, g), + from g.to_continuous_linear_equiv.iterated_fderiv_within_comp_right f hs hx i, + rw [this, continuous_multilinear_map.norm_comp_continuous_linear_isometry_equiv] +end + +/-- Composition with a linear isometry on the right preserves the norm of the iterated derivative +within a set. -/ +lemma linear_isometry_equiv.norm_iterated_fderiv_comp_right + (g : G ≃ₗᵢ[𝕜] E) (f : E → F) (x : G) (i : ℕ) : + ‖iterated_fderiv 𝕜 i (f ∘ g) x‖ = ‖iterated_fderiv 𝕜 i f (g x)‖ := +begin + simp only [← iterated_fderiv_within_univ], + apply g.norm_iterated_fderiv_within_comp_right f unique_diff_on_univ (mem_univ (g x)) i, +end + /-- Composition by continuous linear equivs on the right respects higher differentiability at a point in a domain. -/ lemma continuous_linear_equiv.cont_diff_within_at_comp_iff (e : G ≃L[𝕜] E) : @@ -2028,22 +2193,14 @@ lemma cont_diff_on.comp cont_diff_on 𝕜 n (g ∘ f) s := begin /- we lift all the spaces to a common universe, as we have already proved the result in this - situation. For the lift, we use the trick that `H` is isomorphic through a - continuous linear equiv to `continuous_multilinear_map 𝕜 (λ (i : fin 0), (E × F × G)) H`, and - continuous linear equivs respect smoothness classes. -/ - let Eu := continuous_multilinear_map 𝕜 (λ (i : fin 0), (E × F × G)) E, - letI : normed_add_comm_group Eu := by apply_instance, - letI : normed_space 𝕜 Eu := by apply_instance, - let Fu := continuous_multilinear_map 𝕜 (λ (i : fin 0), (E × F × G)) F, - letI : normed_add_comm_group Fu := by apply_instance, - letI : normed_space 𝕜 Fu := by apply_instance, - let Gu := continuous_multilinear_map 𝕜 (λ (i : fin 0), (E × F × G)) G, - letI : normed_add_comm_group Gu := by apply_instance, - letI : normed_space 𝕜 Gu := by apply_instance, + situation. -/ + let Eu : Type (max uE uF uG) := ulift E, + let Fu : Type (max uE uF uG) := ulift.{(max uE uG) uF} F, + let Gu : Type (max uE uF uG) := ulift.{(max uE uF) uG} G, -- declare the isomorphisms - let isoE : Eu ≃L[𝕜] E := continuous_multilinear_curry_fin0 𝕜 (E × F × G) E, - let isoF : Fu ≃L[𝕜] F := continuous_multilinear_curry_fin0 𝕜 (E × F × G) F, - let isoG : Gu ≃L[𝕜] G := continuous_multilinear_curry_fin0 𝕜 (E × F × G) G, + have isoE : Eu ≃L[𝕜] E := continuous_linear_equiv.ulift, + have isoF : Fu ≃L[𝕜] F := continuous_linear_equiv.ulift, + have isoG : Gu ≃L[𝕜] G := continuous_linear_equiv.ulift, -- lift the functions to the new spaces, check smoothness there, and then go back. let fu : Eu → Fu := (isoF.symm ∘ f) ∘ isoE, have fu_diff : cont_diff_on 𝕜 n fu (isoE ⁻¹' s), @@ -2453,7 +2610,7 @@ cont_diff_within_at.fderiv_within cont_diff_within_at_id hs hmn hx₀s (by rw [preimage_id']) /-- `x ↦ fderiv 𝕜 (f x) (g x)` is smooth at `x₀`. -/ -lemma cont_diff_at.cont_diff_at_fderiv {f : E → F → G} {g : E → F} {n : ℕ∞} +lemma cont_diff_at.fderiv {f : E → F → G} {g : E → F} {n : ℕ∞} (hf : cont_diff_at 𝕜 n (function.uncurry f) (x₀, g x₀)) (hg : cont_diff_at 𝕜 m g x₀) (hmn : m + 1 ≤ n) : @@ -2465,11 +2622,21 @@ begin rw [preimage_univ] end +/-- `fderiv 𝕜 f` is smooth at `x₀`. -/ +lemma cont_diff_at.fderiv_right (hf : cont_diff_at 𝕜 n f x₀) (hmn : (m + 1 : ℕ∞) ≤ n) : + cont_diff_at 𝕜 m (fderiv 𝕜 f) x₀ := +cont_diff_at.fderiv (cont_diff_at.comp (x₀, x₀) hf cont_diff_at_snd) cont_diff_at_id hmn + /-- `x ↦ fderiv 𝕜 (f x) (g x)` is smooth. -/ lemma cont_diff.fderiv {f : E → F → G} {g : E → F} {n m : ℕ∞} (hf : cont_diff 𝕜 m $ function.uncurry f) (hg : cont_diff 𝕜 n g) (hnm : n + 1 ≤ m) : cont_diff 𝕜 n (λ x, fderiv 𝕜 (f x) (g x)) := -cont_diff_iff_cont_diff_at.mpr $ λ x, hf.cont_diff_at.cont_diff_at_fderiv hg.cont_diff_at hnm +cont_diff_iff_cont_diff_at.mpr $ λ x, hf.cont_diff_at.fderiv hg.cont_diff_at hnm + +/-- `fderiv 𝕜 f` is smooth. -/ +lemma cont_diff.fderiv_right (hf : cont_diff 𝕜 n f) (hmn : (m + 1 : ℕ∞) ≤ n) : + cont_diff 𝕜 m (fderiv 𝕜 f) := +cont_diff_iff_cont_diff_at.mpr $ λ x, hf.cont_diff_at.fderiv_right hmn /-- `x ↦ fderiv 𝕜 (f x) (g x)` is continuous. -/ lemma continuous.fderiv {f : E → F → G} {g : E → F} {n : ℕ∞} @@ -2616,6 +2783,9 @@ lemma cont_diff_on.add {s : set E} {f g : E → F} variables {i : ℕ} +/-- The iterated derivative of the sum of two functions is the sum of the iterated derivatives. +See also `iterated_fderiv_within_add_apply'`, which uses the spelling `(λ x, f x + g x)` +instead of `f + g`. -/ lemma iterated_fderiv_within_add_apply {f g : E → F} (hf : cont_diff_on 𝕜 i f s) (hg : cont_diff_on 𝕜 i g s) (hu : unique_diff_on 𝕜 s) (hx : x ∈ s) : @@ -2648,6 +2818,17 @@ begin ... = (iterated_fderiv_within 𝕜 (i+1) f s + iterated_fderiv_within 𝕜 (i+1) g s) x h : rfl } end +/-- The iterated derivative of the sum of two functions is the sum of the iterated derivatives. +This is the same as `iterated_fderiv_within_add_apply`, but using the spelling `(λ x, f x + g x)` +instead of `f + g`, which can be handy for some rewrites. +TODO: use one form consistently. -/ +lemma iterated_fderiv_within_add_apply' {f g : E → F} + (hf : cont_diff_on 𝕜 i f s) (hg : cont_diff_on 𝕜 i g s) (hu : unique_diff_on 𝕜 s) + (hx : x ∈ s) : +iterated_fderiv_within 𝕜 i (λ x, f x + g x) s x = + iterated_fderiv_within 𝕜 i f s x + iterated_fderiv_within 𝕜 i g s x := +iterated_fderiv_within_add_apply hf hg hu hx + lemma iterated_fderiv_add_apply {i : ℕ} {f g : E → F} (hf : cont_diff 𝕜 i f) (hg : cont_diff 𝕜 i g) : iterated_fderiv 𝕜 i (f + g) x = iterated_fderiv 𝕜 i f x + iterated_fderiv 𝕜 i g x := @@ -2656,6 +2837,11 @@ begin exact iterated_fderiv_within_add_apply hf hg unique_diff_on_univ (set.mem_univ _), end +lemma iterated_fderiv_add_apply' {i : ℕ} {f g : E → F} (hf : cont_diff 𝕜 i f) + (hg : cont_diff 𝕜 i g) : + iterated_fderiv 𝕜 i (λ x, f x + g x) x = iterated_fderiv 𝕜 i f x + iterated_fderiv 𝕜 i g x := +iterated_fderiv_add_apply hf hg + end add /-! ### Negative -/ diff --git a/src/analysis/calculus/fderiv.lean b/src/analysis/calculus/fderiv.lean index a5b0412a771bb..5953792f5af23 100644 --- a/src/analysis/calculus/fderiv.lean +++ b/src/analysis/calculus/fderiv.lean @@ -851,7 +851,7 @@ lemma filter.eventually_eq.fderiv_within_eq_nhds (hs : unique_diff_within_at (show f₁ =ᶠ[𝓝[s] x] f, from nhds_within_le_nhds hL).fderiv_within_eq hs (mem_of_mem_nhds hL : _) lemma fderiv_within_congr (hs : unique_diff_within_at 𝕜 s x) - (hL : ∀y∈s, f₁ y = f y) (hx : f₁ x = f x) : + (hL : ∀ y ∈ s, f₁ y = f y) (hx : f₁ x = f x) : fderiv_within 𝕜 f₁ s x = fderiv_within 𝕜 f s x := begin apply filter.eventually_eq.fderiv_within_eq hs _ hx, @@ -859,6 +859,11 @@ begin exact hL end +lemma fderiv_within_congr' (hs : unique_diff_within_at 𝕜 s x) + (hL : ∀ y ∈ s, f₁ y = f y) (hx : x ∈ s) : + fderiv_within 𝕜 f₁ s x = fderiv_within 𝕜 f s x := +fderiv_within_congr hs hL (hL x hx) + lemma filter.eventually_eq.fderiv_eq (hL : f₁ =ᶠ[𝓝 x] f) : fderiv 𝕜 f₁ x = fderiv 𝕜 f x := begin @@ -2162,7 +2167,7 @@ end sub section bilinear_map /-! ### Derivative of a bounded bilinear map -/ -variables {b : E × F → G} {u : set (E × F) } +variables {b : E × F → G} {u : set (E × F)} open normed_field @@ -2228,6 +2233,35 @@ lemma is_bounded_bilinear_map.differentiable_on (h : is_bounded_bilinear_map differentiable_on 𝕜 b u := h.differentiable.differentiable_on +variable (B : E →L[𝕜] F →L[𝕜] G) + +lemma continuous_linear_map.has_fderiv_within_at_of_bilinear + {f : G' → E} {g : G' → F} {f' : G' →L[𝕜] E} {g' : G' →L[𝕜] F} {x : G'} {s : set G'} + (hf : has_fderiv_within_at f f' s x) (hg : has_fderiv_within_at g g' s x) : + has_fderiv_within_at (λ y, B (f y) (g y)) (B.precompR G' (f x) g' + B.precompL G' f' (g x)) s x := +(B.is_bounded_bilinear_map.has_fderiv_at (f x, g x)).comp_has_fderiv_within_at x (hf.prod hg) + +lemma continuous_linear_map.has_fderiv_at_of_bilinear + {f : G' → E} {g : G' → F} {f' : G' →L[𝕜] E} {g' : G' →L[𝕜] F} {x : G'} + (hf : has_fderiv_at f f' x) (hg : has_fderiv_at g g' x) : + has_fderiv_at (λ y, B (f y) (g y)) (B.precompR G' (f x) g' + B.precompL G' f' (g x)) x := +(B.is_bounded_bilinear_map.has_fderiv_at (f x, g x)).comp x (hf.prod hg) + +lemma continuous_linear_map.fderiv_within_of_bilinear + {f : G' → E} {g : G' → F} {x : G'} {s : set G'} + (hf : differentiable_within_at 𝕜 f s x) (hg : differentiable_within_at 𝕜 g s x) + (hs : unique_diff_within_at 𝕜 s x) : + fderiv_within 𝕜 (λ y, B (f y) (g y)) s x = + (B.precompR G' (f x) (fderiv_within 𝕜 g s x) + B.precompL G' (fderiv_within 𝕜 f s x) (g x)) := +(B.has_fderiv_within_at_of_bilinear hf.has_fderiv_within_at hg.has_fderiv_within_at).fderiv_within + hs + +lemma continuous_linear_map.fderiv_of_bilinear {f : G' → E} {g : G' → F} {x : G'} + (hf : differentiable_at 𝕜 f x) (hg : differentiable_at 𝕜 g x) : + fderiv 𝕜 (λ y, B (f y) (g y)) x = + (B.precompR G' (f x) (fderiv 𝕜 g x) + B.precompL G' (fderiv 𝕜 f x) (g x)) := +(B.has_fderiv_at_of_bilinear hf.has_fderiv_at hg.has_fderiv_at).fderiv + end bilinear_map section clm_comp_apply @@ -2772,6 +2806,88 @@ begin exact iso.comp_fderiv_within unique_diff_within_at_univ, end +lemma comp_right_differentiable_within_at_iff {f : F → G} {s : set F} {x : E} : + differentiable_within_at 𝕜 (f ∘ iso) (iso ⁻¹' s) x ↔ differentiable_within_at 𝕜 f s (iso x) := +begin + refine ⟨λ H, _, λ H, H.comp x iso.differentiable_within_at (maps_to_preimage _ s)⟩, + have : differentiable_within_at 𝕜 ((f ∘ iso) ∘ iso.symm) s (iso x), + { rw ← iso.symm_apply_apply x at H, + apply H.comp (iso x) iso.symm.differentiable_within_at, + assume y hy, + simpa only [mem_preimage, apply_symm_apply] using hy }, + rwa [function.comp.assoc, iso.self_comp_symm] at this, +end + +lemma comp_right_differentiable_at_iff {f : F → G} {x : E} : + differentiable_at 𝕜 (f ∘ iso) x ↔ differentiable_at 𝕜 f (iso x) := +by simp only [← differentiable_within_at_univ, ← iso.comp_right_differentiable_within_at_iff, + preimage_univ] + +lemma comp_right_differentiable_on_iff {f : F → G} {s : set F} : + differentiable_on 𝕜 (f ∘ iso) (iso ⁻¹' s) ↔ differentiable_on 𝕜 f s := +begin + refine ⟨λ H y hy, _, λ H y hy, iso.comp_right_differentiable_within_at_iff.2 (H _ hy)⟩, + rw [← iso.apply_symm_apply y, ← comp_right_differentiable_within_at_iff], + apply H, + simpa only [mem_preimage, apply_symm_apply] using hy, +end + +lemma comp_right_differentiable_iff {f : F → G} : + differentiable 𝕜 (f ∘ iso) ↔ differentiable 𝕜 f := +by simp only [← differentiable_on_univ, ← iso.comp_right_differentiable_on_iff, preimage_univ] + +lemma comp_right_has_fderiv_within_at_iff + {f : F → G} {s : set F} {x : E} {f' : F →L[𝕜] G} : + has_fderiv_within_at (f ∘ iso) (f'.comp (iso : E →L[𝕜] F)) (iso ⁻¹' s) x ↔ + has_fderiv_within_at f f' s (iso x) := +begin + refine ⟨λ H, _, λ H, H.comp x iso.has_fderiv_within_at (maps_to_preimage _ s)⟩, + rw [← iso.symm_apply_apply x] at H, + have A : f = (f ∘ iso) ∘ iso.symm, by { rw [function.comp.assoc, iso.self_comp_symm], refl }, + have B : f' = (f'.comp (iso : E →L[𝕜] F)).comp (iso.symm : F →L[𝕜] E), + by rw [continuous_linear_map.comp_assoc, iso.coe_comp_coe_symm, + continuous_linear_map.comp_id], + rw [A, B], + apply H.comp (iso x) iso.symm.has_fderiv_within_at, + assume y hy, + simpa only [mem_preimage, apply_symm_apply] using hy +end + +lemma comp_right_has_fderiv_at_iff {f : F → G} {x : E} {f' : F →L[𝕜] G} : + has_fderiv_at (f ∘ iso) (f'.comp (iso : E →L[𝕜] F)) x ↔ has_fderiv_at f f' (iso x) := +by simp only [← has_fderiv_within_at_univ, ← comp_right_has_fderiv_within_at_iff, preimage_univ] + +lemma comp_right_has_fderiv_within_at_iff' + {f : F → G} {s : set F} {x : E} {f' : E →L[𝕜] G} : + has_fderiv_within_at (f ∘ iso) f' (iso ⁻¹' s) x ↔ + has_fderiv_within_at f (f'.comp (iso.symm : F →L[𝕜] E)) s (iso x) := +by rw [← iso.comp_right_has_fderiv_within_at_iff, continuous_linear_map.comp_assoc, + iso.coe_symm_comp_coe, continuous_linear_map.comp_id] + +lemma comp_right_has_fderiv_at_iff' {f : F → G} {x : E} {f' : E →L[𝕜] G} : + has_fderiv_at (f ∘ iso) f' x ↔ has_fderiv_at f (f'.comp (iso.symm : F →L[𝕜] E)) (iso x) := +by simp only [← has_fderiv_within_at_univ, ← iso.comp_right_has_fderiv_within_at_iff', + preimage_univ] + +lemma comp_right_fderiv_within {f : F → G} {s : set F} {x : E} + (hxs : unique_diff_within_at 𝕜 (iso ⁻¹' s) x) : + fderiv_within 𝕜 (f ∘ iso) (iso ⁻¹'s) x = (fderiv_within 𝕜 f s (iso x)).comp (iso : E →L[𝕜] F) := +begin + by_cases h : differentiable_within_at 𝕜 f s (iso x), + { exact (iso.comp_right_has_fderiv_within_at_iff.2 (h.has_fderiv_within_at)).fderiv_within hxs }, + { have : ¬ differentiable_within_at 𝕜 (f ∘ iso) (iso ⁻¹' s) x, + { assume h', exact h (iso.comp_right_differentiable_within_at_iff.1 h') }, + rw [fderiv_within_zero_of_not_differentiable_within_at h, + fderiv_within_zero_of_not_differentiable_within_at this, continuous_linear_map.zero_comp] } +end + +lemma comp_right_fderiv {f : F → G} {x : E} : + fderiv 𝕜 (f ∘ iso) x = (fderiv 𝕜 f (iso x)).comp (iso : E →L[𝕜] F) := +begin + rw [← fderiv_within_univ, ← fderiv_within_univ, ← iso.comp_right_fderiv_within, preimage_univ], + exact unique_diff_within_at_univ, +end + end continuous_linear_equiv namespace linear_isometry_equiv diff --git a/src/analysis/normed_space/linear_isometry.lean b/src/analysis/normed_space/linear_isometry.lean index 95db39fb82021..b4088ddb2268a 100644 --- a/src/analysis/normed_space/linear_isometry.lean +++ b/src/analysis/normed_space/linear_isometry.lean @@ -514,6 +514,11 @@ variables (R E) /-- Identity map as a `linear_isometry_equiv`. -/ def refl : E ≃ₗᵢ[R] E := ⟨linear_equiv.refl R E, λ x, rfl⟩ +/-- Linear isometry equiv between a space and its lift to another universe. -/ +def ulift : ulift E ≃ₗᵢ[R] E := +{ norm_map' := λ x, rfl, + .. continuous_linear_equiv.ulift } + variables {R E} instance : inhabited (E ≃ₗᵢ[R] E) := ⟨refl R E⟩ diff --git a/src/analysis/normed_space/multilinear.lean b/src/analysis/normed_space/multilinear.lean index 139ff4b8529c4..fab1ad5232064 100644 --- a/src/analysis/normed_space/multilinear.lean +++ b/src/analysis/normed_space/multilinear.lean @@ -874,7 +874,45 @@ linear_map.mk_continuous₂ (λ f g₁ g₂, by { ext1, apply f.map_add }) (λ c f g, by { ext1, simp })) 1 $ λ f g, by { rw one_mul, exact f.norm_comp_continuous_multilinear_map_le g } -variables {𝕜 E G G'} +variables {𝕜 G G'} + +/-- `continuous_linear_map.comp_continuous_multilinear_map` as a bundled +continuous linear equiv. -/ +def _root_.continuous_linear_equiv.comp_continuous_multilinear_mapL (g : G ≃L[𝕜] G') : + continuous_multilinear_map 𝕜 E G ≃L[𝕜] continuous_multilinear_map 𝕜 E G' := +{ inv_fun := comp_continuous_multilinear_mapL 𝕜 _ _ _ g.symm.to_continuous_linear_map, + left_inv := begin + assume f, + ext1 m, + simp only [comp_continuous_multilinear_mapL, continuous_linear_equiv.coe_def_rev, + to_linear_map_eq_coe, linear_map.to_fun_eq_coe, coe_coe, linear_map.mk_continuous₂_apply, + linear_map.mk₂_apply, comp_continuous_multilinear_map_coe, continuous_linear_equiv.coe_coe, + function.comp_app, continuous_linear_equiv.symm_apply_apply], + end, + right_inv := begin + assume f, + ext1 m, + simp only [comp_continuous_multilinear_mapL, continuous_linear_equiv.coe_def_rev, + to_linear_map_eq_coe, linear_map.mk_continuous₂_apply, linear_map.mk₂_apply, + linear_map.to_fun_eq_coe, coe_coe, comp_continuous_multilinear_map_coe, + continuous_linear_equiv.coe_coe, function.comp_app, continuous_linear_equiv.apply_symm_apply], + end, + continuous_to_fun := + (comp_continuous_multilinear_mapL 𝕜 _ _ _ g.to_continuous_linear_map).continuous, + continuous_inv_fun := + (comp_continuous_multilinear_mapL 𝕜 _ _ _ g.symm.to_continuous_linear_map).continuous, + .. comp_continuous_multilinear_mapL 𝕜 _ _ _ g.to_continuous_linear_map } + +@[simp] lemma _root_.continuous_linear_equiv.comp_continuous_multilinear_mapL_symm + (g : G ≃L[𝕜] G') : + (g.comp_continuous_multilinear_mapL E).symm = g.symm.comp_continuous_multilinear_mapL E := rfl + +variables {E} + +@[simp] lemma _root_.continuous_linear_equiv.comp_continuous_multilinear_mapL_apply + (g : G ≃L[𝕜] G') (f : continuous_multilinear_map 𝕜 E G) : + g.comp_continuous_multilinear_mapL E f = (g : G →L[𝕜] G').comp_continuous_multilinear_map f := +rfl /-- Flip arguments in `f : G →L[𝕜] continuous_multilinear_map 𝕜 E G'` to get `continuous_multilinear_map 𝕜 E (G →L[𝕜] G')` -/ @@ -899,6 +937,14 @@ multilinear_map.mk_continuous (mul_nonneg (norm_nonneg f) (prod_nonneg $ λ i hi, norm_nonneg (m i))) _ end continuous_linear_map + +lemma linear_isometry.norm_comp_continuous_multilinear_map + (g : G →ₗᵢ[𝕜] G') (f : continuous_multilinear_map 𝕜 E G) : + ‖g.to_continuous_linear_map.comp_continuous_multilinear_map f‖ = ‖f‖ := +by simp only [continuous_linear_map.comp_continuous_multilinear_map_coe, + linear_isometry.coe_to_continuous_linear_map, linear_isometry.norm_map, + continuous_multilinear_map.norm_def] + open continuous_multilinear_map namespace multilinear_map @@ -979,6 +1025,31 @@ calc ‖g (λ i, f i (m i))‖ ≤ ‖g‖ * ∏ i, ‖f i (m i)‖ : g.le_op_no (prod_le_prod (λ _ _, norm_nonneg _) (λ i hi, (f i).le_op_norm (m i))) (norm_nonneg g) ... = (‖g‖ * ∏ i, ‖f i‖) * ∏ i, ‖m i‖ : by rw [prod_mul_distrib, mul_assoc] +lemma norm_comp_continuous_linear_isometry_le (g : continuous_multilinear_map 𝕜 E₁ G) + (f : Π i, E i →ₗᵢ[𝕜] E₁ i) : + ‖g.comp_continuous_linear_map (λ i, (f i).to_continuous_linear_map)‖ ≤ ‖g‖ := +begin + apply op_norm_le_bound _ (norm_nonneg _) (λ m, _), + apply (g.le_op_norm _).trans _, + simp only [continuous_linear_map.to_linear_map_eq_coe, continuous_linear_map.coe_coe, + linear_isometry.coe_to_continuous_linear_map, linear_isometry.norm_map] +end + +lemma norm_comp_continuous_linear_isometry_equiv (g : continuous_multilinear_map 𝕜 E₁ G) + (f : Π i, E i ≃ₗᵢ[𝕜] E₁ i) : + ‖g.comp_continuous_linear_map (λ i, (f i : E i →L[𝕜] E₁ i))‖ = ‖g‖ := +begin + apply le_antisymm (g.norm_comp_continuous_linear_isometry_le (λ i, (f i).to_linear_isometry)), + have : g = (g.comp_continuous_linear_map (λ i, (f i : E i →L[𝕜] E₁ i))) + .comp_continuous_linear_map (λ i, ((f i).symm : E₁ i →L[𝕜] E i)), + { ext1 m, + simp only [comp_continuous_linear_map_apply, linear_isometry_equiv.coe_coe'', + linear_isometry_equiv.apply_symm_apply] }, + conv_lhs { rw this }, + apply (g.comp_continuous_linear_map (λ i, (f i : E i →L[𝕜] E₁ i))) + .norm_comp_continuous_linear_isometry_le (λ i, (f i).symm.to_linear_isometry), +end + /-- `continuous_multilinear_map.comp_continuous_linear_map` as a bundled continuous linear map. This implementation fixes `f : Π i, E i →L[𝕜] E₁ i`. @@ -992,8 +1063,8 @@ linear_map.mk_continuous map_smul' := λ c g, rfl } (∏ i, ‖f i‖) $ λ g, (norm_comp_continuous_linear_le _ _).trans_eq (mul_comm _ _) -@[simp] lemma comp_continuous_linear_mapL_apply (g : continuous_multilinear_map 𝕜 E₁ G) - (f : Π i, E i →L[𝕜] E₁ i) : +@[simp] lemma comp_continuous_linear_mapL_apply + (g : continuous_multilinear_map 𝕜 E₁ G) (f : Π i, E i →L[𝕜] E₁ i) : comp_continuous_linear_mapL f g = g.comp_continuous_linear_map f := rfl @@ -1001,6 +1072,45 @@ lemma norm_comp_continuous_linear_mapL_le (f : Π i, E i →L[𝕜] E₁ i) : ‖@comp_continuous_linear_mapL 𝕜 ι E E₁ G _ _ _ _ _ _ _ _ _ f‖ ≤ (∏ i, ‖f i‖) := linear_map.mk_continuous_norm_le _ (prod_nonneg $ λ i _, norm_nonneg _) _ +variable (G) + +/-- `continuous_multilinear_map.comp_continuous_linear_map` as a bundled continuous linear equiv, +given `f : Π i, E i ≃L[𝕜] E₁ i`. -/ +def comp_continuous_linear_map_equivL (f : Π i, E i ≃L[𝕜] E₁ i) : + continuous_multilinear_map 𝕜 E₁ G ≃L[𝕜] continuous_multilinear_map 𝕜 E G := +{ inv_fun := comp_continuous_linear_mapL (λ i, ((f i).symm : E₁ i →L[𝕜] E i)), + continuous_to_fun := (comp_continuous_linear_mapL (λ i, (f i : E i →L[𝕜] E₁ i))).continuous, + continuous_inv_fun := + (comp_continuous_linear_mapL (λ i, ((f i).symm : E₁ i →L[𝕜] E i))).continuous, + left_inv := begin + assume g, + ext1 m, + simp only [continuous_linear_map.to_linear_map_eq_coe, linear_map.to_fun_eq_coe, + continuous_linear_map.coe_coe, comp_continuous_linear_mapL_apply, + comp_continuous_linear_map_apply, continuous_linear_equiv.coe_coe, + continuous_linear_equiv.apply_symm_apply], + end, + right_inv := begin + assume g, + ext1 m, + simp only [continuous_linear_map.to_linear_map_eq_coe, comp_continuous_linear_mapL_apply, + linear_map.to_fun_eq_coe, continuous_linear_map.coe_coe, comp_continuous_linear_map_apply, + continuous_linear_equiv.coe_coe, continuous_linear_equiv.symm_apply_apply], + end, + .. comp_continuous_linear_mapL (λ i, (f i : E i →L[𝕜] E₁ i)) } + +@[simp] lemma comp_continuous_linear_map_equivL_symm (f : Π i, E i ≃L[𝕜] E₁ i) : + (comp_continuous_linear_map_equivL G f).symm = + comp_continuous_linear_map_equivL G (λ (i : ι), (f i).symm) := +rfl + +variable {G} + +@[simp] lemma comp_continuous_linear_map_equivL_apply + (g : continuous_multilinear_map 𝕜 E₁ G) (f : Π i, E i ≃L[𝕜] E₁ i) : + comp_continuous_linear_map_equivL G f g = + g.comp_continuous_linear_map (λ i, (f i : E i →L[𝕜] E₁ i)) := rfl + end continuous_multilinear_map section smul diff --git a/src/analysis/normed_space/operator_norm.lean b/src/analysis/normed_space/operator_norm.lean index d346e8da0b119..b710ee3a5e061 100644 --- a/src/analysis/normed_space/operator_norm.lean +++ b/src/analysis/normed_space/operator_norm.lean @@ -789,10 +789,13 @@ linear_map.mk_continuous₂ function.comp_app, pi.smul_apply] })) 1 $ λ f g, by simpa only [one_mul] using op_norm_comp_le f g -variables {𝕜 σ₁₂ σ₂₃ E F G} - include σ₁₃ +lemma norm_compSL_le : ‖compSL E F G σ₁₂ σ₂₃‖ ≤ 1 := +linear_map.mk_continuous₂_norm_le _ zero_le_one _ + +variables {𝕜 σ₁₂ σ₂₃ E F G} + @[simp] lemma compSL_apply (f : F →SL[σ₂₃] G) (g : E →SL[σ₁₂] F) : compSL E F G σ₁₂ σ₂₃ f g = f.comp g := rfl @@ -811,7 +814,10 @@ variables (𝕜 σ₁₂ σ₂₃ E Fₗ Gₗ) /-- Composition of continuous linear maps as a continuous bilinear map. -/ def compL : (Fₗ →L[𝕜] Gₗ) →L[𝕜] (E →L[𝕜] Fₗ) →L[𝕜] (E →L[𝕜] Gₗ) := - compSL E Fₗ Gₗ (ring_hom.id 𝕜) (ring_hom.id 𝕜) +compSL E Fₗ Gₗ (ring_hom.id 𝕜) (ring_hom.id 𝕜) + +lemma norm_compL_le : ‖compL 𝕜 E Fₗ Gₗ‖ ≤ 1 := +norm_compSL_le _ _ _ _ _ @[simp] lemma compL_apply (f : Fₗ →L[𝕜] Gₗ) (g : E →L[𝕜] Fₗ) : compL 𝕜 E Fₗ Gₗ f g = f.comp g := rfl @@ -825,6 +831,14 @@ def precompR (L : E →L[𝕜] Fₗ →L[𝕜] Gₗ) : E →L[𝕜] (Eₗ →L[ def precompL (L : E →L[𝕜] Fₗ →L[𝕜] Gₗ) : (Eₗ →L[𝕜] E) →L[𝕜] Fₗ →L[𝕜] (Eₗ →L[𝕜] Gₗ) := (precompR Eₗ (flip L)).flip +lemma norm_precompR_le (L : E →L[𝕜] Fₗ →L[𝕜] Gₗ) : ‖precompR Eₗ L‖ ≤ ‖L‖ := calc +‖precompR Eₗ L‖ ≤ ‖compL 𝕜 Eₗ Fₗ Gₗ‖ * ‖L‖ : op_norm_comp_le _ _ +... ≤ 1 * ‖L‖ : mul_le_mul_of_nonneg_right (norm_compL_le _ _ _ _) (norm_nonneg _) +... = ‖L‖ : by rw one_mul + +lemma norm_precompL_le (L : E →L[𝕜] Fₗ →L[𝕜] Gₗ) : ‖precompL Eₗ L‖ ≤ ‖L‖ := +by { rw [precompL, op_norm_flip, ← op_norm_flip L], exact norm_precompR_le _ L.flip } + section prod universes u₁ u₂ u₃ u₄ diff --git a/src/topology/algebra/module/multilinear.lean b/src/topology/algebra/module/multilinear.lean index 3d3976403f62b..5cc8f9997f417 100644 --- a/src/topology/algebra/module/multilinear.lean +++ b/src/topology/algebra/module/multilinear.lean @@ -256,7 +256,6 @@ def _root_.continuous_linear_map.comp_continuous_multilinear_map (g : M₂ → M₃) ∘ (f : (Πi, M₁ i) → M₂) := by { ext m, refl } - /-- `continuous_multilinear_map.pi` as an `equiv`. -/ @[simps] def pi_equiv {ι' : Type*} {M' : ι' → Type*} [Π i, add_comm_monoid (M' i)] From aa3a420527e0fbfd0f6615b95b761254a9166e12 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Tue, 21 Mar 2023 17:28:57 +0000 Subject: [PATCH 0668/1291] chore(*): add mathlib4 synchronization comments (#18625) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.direct_sum.finsupp` * `algebra.free_algebra` * `ring_theory.noetherian` --- src/algebra/direct_sum/finsupp.lean | 3 +++ src/algebra/free_algebra.lean | 3 +++ src/ring_theory/noetherian.lean | 3 +++ 3 files changed, 9 insertions(+) diff --git a/src/algebra/direct_sum/finsupp.lean b/src/algebra/direct_sum/finsupp.lean index 279482cfa0cb2..e5ba9086e4008 100644 --- a/src/algebra/direct_sum/finsupp.lean +++ b/src/algebra/direct_sum/finsupp.lean @@ -9,6 +9,9 @@ import data.finsupp.to_dfinsupp /-! # Results on direct sums and finitely supported functions. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + 1. The linear equivalence between finitely supported functions `ι →₀ M` and the direct sum of copies of `M` indexed by `ι`. -/ diff --git a/src/algebra/free_algebra.lean b/src/algebra/free_algebra.lean index 965190730dab2..d3e5a6ebcd813 100644 --- a/src/algebra/free_algebra.lean +++ b/src/algebra/free_algebra.lean @@ -9,6 +9,9 @@ import algebra.monoid_algebra.basic /-! # Free Algebras +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given a commutative semiring `R`, and a type `X`, we construct the free unital, associative `R`-algebra on `X`. diff --git a/src/ring_theory/noetherian.lean b/src/ring_theory/noetherian.lean index 612de749225cf..d2110c7afeeca 100644 --- a/src/ring_theory/noetherian.lean +++ b/src/ring_theory/noetherian.lean @@ -16,6 +16,9 @@ import ring_theory.nilpotent /-! # Noetherian rings and modules +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The following are equivalent for a module M over a ring R: 1. Every increasing chain of submodules M₁ ⊆ M₂ ⊆ M₃ ⊆ ⋯ eventually stabilises. 2. Every submodule is finitely generated. From dd6388c44e6f6b4547070b887c5905d5cfe6c9f8 Mon Sep 17 00:00:00 2001 From: sgouezel Date: Tue, 21 Mar 2023 19:41:17 +0000 Subject: [PATCH 0669/1291] chore(analysis/calculus/cont_diff): split into two files (#18627) The file `cont_diff.lean` takes 3824 lines. We cut it into two files, `cont_diff_def.lean` and `cont_diff.lean`, the former (of 1597 lines) containing general definitions and the latter (of 2264 lines) containing results on the preservation of smoothness by the usual operations and various API extensions that depend on these. There is no single change in maths or code, this is purely a split PR. --- src/analysis/calculus/cont_diff.lean | 1570 +-------------------- src/analysis/calculus/cont_diff_def.lean | 1595 ++++++++++++++++++++++ 2 files changed, 1601 insertions(+), 1564 deletions(-) create mode 100644 src/analysis/calculus/cont_diff_def.lean diff --git a/src/analysis/calculus/cont_diff.lean b/src/analysis/calculus/cont_diff.lean index 37d469b3cc08a..d4518001e7b6d 100644 --- a/src/analysis/calculus/cont_diff.lean +++ b/src/analysis/calculus/cont_diff.lean @@ -3,149 +3,15 @@ Copyright (c) 2019 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ +import analysis.calculus.cont_diff_def import analysis.calculus.mean_value import analysis.normed_space.finite_dimension -import analysis.normed_space.multilinear -import analysis.calculus.formal_multilinear_series -import data.enat.basic -import tactic.congrm /-! -# Higher differentiability - -A function is `C^1` on a domain if it is differentiable there, and its derivative is continuous. -By induction, it is `C^n` if it is `C^{n-1}` and its (n-1)-th derivative is `C^1` there or, -equivalently, if it is `C^1` and its derivative is `C^{n-1}`. -Finally, it is `C^∞` if it is `C^n` for all n. - -We formalize these notions by defining iteratively the `n+1`-th derivative of a function as the -derivative of the `n`-th derivative. It is called `iterated_fderiv 𝕜 n f x` where `𝕜` is the -field, `n` is the number of iterations, `f` is the function and `x` is the point, and it is given -as an `n`-multilinear map. We also define a version `iterated_fderiv_within` relative to a domain, -as well as predicates `cont_diff_within_at`, `cont_diff_at`, `cont_diff_on` and -`cont_diff` saying that the function is `C^n` within a set at a point, at a point, on a set -and on the whole space respectively. - -To avoid the issue of choice when choosing a derivative in sets where the derivative is not -necessarily unique, `cont_diff_on` is not defined directly in terms of the -regularity of the specific choice `iterated_fderiv_within 𝕜 n f s` inside `s`, but in terms of the -existence of a nice sequence of derivatives, expressed with a predicate -`has_ftaylor_series_up_to_on`. - -We prove basic properties of these notions. - -## Main definitions and results -Let `f : E → F` be a map between normed vector spaces over a nontrivially normed field `𝕜`. - -* `has_ftaylor_series_up_to n f p`: expresses that the formal multilinear series `p` is a sequence - of iterated derivatives of `f`, up to the `n`-th term (where `n` is a natural number or `∞`). -* `has_ftaylor_series_up_to_on n f p s`: same thing, but inside a set `s`. The notion of derivative - is now taken inside `s`. In particular, derivatives don't have to be unique. -* `cont_diff 𝕜 n f`: expresses that `f` is `C^n`, i.e., it admits a Taylor series up to - rank `n`. -* `cont_diff_on 𝕜 n f s`: expresses that `f` is `C^n` in `s`. -* `cont_diff_at 𝕜 n f x`: expresses that `f` is `C^n` around `x`. -* `cont_diff_within_at 𝕜 n f s x`: expresses that `f` is `C^n` around `x` within the set `s`. -* `iterated_fderiv_within 𝕜 n f s x` is an `n`-th derivative of `f` over the field `𝕜` on the - set `s` at the point `x`. It is a continuous multilinear map from `E^n` to `F`, defined as a - derivative within `s` of `iterated_fderiv_within 𝕜 (n-1) f s` if one exists, and `0` otherwise. -* `iterated_fderiv 𝕜 n f x` is the `n`-th derivative of `f` over the field `𝕜` at the point `x`. - It is a continuous multilinear map from `E^n` to `F`, defined as a derivative of - `iterated_fderiv 𝕜 (n-1) f` if one exists, and `0` otherwise. - -In sets of unique differentiability, `cont_diff_on 𝕜 n f s` can be expressed in terms of the -properties of `iterated_fderiv_within 𝕜 m f s` for `m ≤ n`. In the whole space, -`cont_diff 𝕜 n f` can be expressed in terms of the properties of `iterated_fderiv 𝕜 m f` -for `m ≤ n`. - -We also prove that the usual operations (addition, multiplication, difference, composition, and -so on) preserve `C^n` functions. - -## Implementation notes - -The definitions in this file are designed to work on any field `𝕜`. They are sometimes slightly more -complicated than the naive definitions one would guess from the intuition over the real or complex -numbers, but they are designed to circumvent the lack of gluing properties and partitions of unity -in general. In the usual situations, they coincide with the usual definitions. - -### Definition of `C^n` functions in domains - -One could define `C^n` functions in a domain `s` by fixing an arbitrary choice of derivatives (this -is what we do with `iterated_fderiv_within`) and requiring that all these derivatives up to `n` are -continuous. If the derivative is not unique, this could lead to strange behavior like two `C^n` -functions `f` and `g` on `s` whose sum is not `C^n`. A better definition is thus to say that a -function is `C^n` inside `s` if it admits a sequence of derivatives up to `n` inside `s`. - -This definition still has the problem that a function which is locally `C^n` would not need to -be `C^n`, as different choices of sequences of derivatives around different points might possibly -not be glued together to give a globally defined sequence of derivatives. (Note that this issue -can not happen over reals, thanks to partition of unity, but the behavior over a general field is -not so clear, and we want a definition for general fields). Also, there are locality -problems for the order parameter: one could image a function which, for each `n`, has a nice -sequence of derivatives up to order `n`, but they do not coincide for varying `n` and can therefore -not be glued to give rise to an infinite sequence of derivatives. This would give a function -which is `C^n` for all `n`, but not `C^∞`. We solve this issue by putting locality conditions -in space and order in our definition of `cont_diff_within_at` and `cont_diff_on`. -The resulting definition is slightly more complicated to work with (in fact not so much), but it -gives rise to completely satisfactory theorems. - -For instance, with this definition, a real function which is `C^m` (but not better) on `(-1/m, 1/m)` -for each natural `m` is by definition `C^∞` at `0`. - -There is another issue with the definition of `cont_diff_within_at 𝕜 n f s x`. We can -require the existence and good behavior of derivatives up to order `n` on a neighborhood of `x` -within `s`. However, this does not imply continuity or differentiability within `s` of the function -at `x` when `x` does not belong to `s`. Therefore, we require such existence and good behavior on -a neighborhood of `x` within `s ∪ {x}` (which appears as `insert x s` in this file). - -### Side of the composition, and universe issues - -With a naïve direct definition, the `n`-th derivative of a function belongs to the space -`E →L[𝕜] (E →L[𝕜] (E ... F)...)))` where there are n iterations of `E →L[𝕜]`. This space -may also be seen as the space of continuous multilinear functions on `n` copies of `E` with -values in `F`, by uncurrying. This is the point of view that is usually adopted in textbooks, -and that we also use. This means that the definition and the first proofs are slightly involved, -as one has to keep track of the uncurrying operation. The uncurrying can be done from the -left or from the right, amounting to defining the `n+1`-th derivative either as the derivative of -the `n`-th derivative, or as the `n`-th derivative of the derivative. -For proofs, it would be more convenient to use the latter approach (from the right), -as it means to prove things at the `n+1`-th step we only need to understand well enough the -derivative in `E →L[𝕜] F` (contrary to the approach from the left, where one would need to know -enough on the `n`-th derivative to deduce things on the `n+1`-th derivative). - -However, the definition from the right leads to a universe polymorphism problem: if we define -`iterated_fderiv 𝕜 (n + 1) f x = iterated_fderiv 𝕜 n (fderiv 𝕜 f) x` by induction, we need to -generalize over all spaces (as `f` and `fderiv 𝕜 f` don't take values in the same space). It is -only possible to generalize over all spaces in some fixed universe in an inductive definition. -For `f : E → F`, then `fderiv 𝕜 f` is a map `E → (E →L[𝕜] F)`. Therefore, the definition will only -work if `F` and `E →L[𝕜] F` are in the same universe. - -This issue does not appear with the definition from the left, where one does not need to generalize -over all spaces. Therefore, we use the definition from the left. This means some proofs later on -become a little bit more complicated: to prove that a function is `C^n`, the most efficient approach -is to exhibit a formula for its `n`-th derivative and prove it is continuous (contrary to the -inductive approach where one would prove smoothness statements without giving a formula for the -derivative). In the end, this approach is still satisfactory as it is good to have formulas for the -iterated derivatives in various constructions. - -One point where we depart from this explicit approach is in the proof of smoothness of a -composition: there is a formula for the `n`-th derivative of a composition (Faà di Bruno's formula), -but it is very complicated and barely usable, while the inductive proof is very simple. Thus, we -give the inductive proof. As explained above, it works by generalizing over the target space, hence -it only works well if all spaces belong to the same universe. To get the general version, we lift -things to a common universe using a trick. - -### Variables management - -The textbook definitions and proofs use various identifications and abuse of notations, for instance -when saying that the natural space in which the derivative lives, i.e., -`E →L[𝕜] (E →L[𝕜] ( ... →L[𝕜] F))`, is the same as a space of multilinear maps. When doing things -formally, we need to provide explicit maps for these identifications, and chase some diagrams to see -everything is compatible with the identifications. In particular, one needs to check that taking the -derivative and then doing the identification, or first doing the identification and then taking the -derivative, gives the same result. The key point for this is that taking the derivative commutes -with continuous linear equivalences. Therefore, we need to implement all our identifications with -continuous linear equivs. +# Higher differentiability of usual operations + +We prove that the usual operations (addition, multiplication, difference, composition, and +so on) preserve `C^n` functions. We also expand the API aound `C^n` functions. ## Notations @@ -178,1431 +44,7 @@ variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] {G : Type uG} [normed_add_comm_group G] [normed_space 𝕜 G] {X : Type*} [normed_add_comm_group X] [normed_space 𝕜 X] {s s₁ t u : set E} {f f₁ : E → F} {g : F → G} {x x₀ : E} {c : F} -{b : E × F → G} {m n : ℕ∞} - -/-! ### Functions with a Taylor series on a domain -/ - -variable {p : E → formal_multilinear_series 𝕜 E F} - -/-- `has_ftaylor_series_up_to_on n f p s` registers the fact that `p 0 = f` and `p (m+1)` is a -derivative of `p m` for `m < n`, and is continuous for `m ≤ n`. This is a predicate analogous to -`has_fderiv_within_at` but for higher order derivatives. -/ -structure has_ftaylor_series_up_to_on (n : ℕ∞) - (f : E → F) (p : E → formal_multilinear_series 𝕜 E F) (s : set E) : Prop := -(zero_eq : ∀ x ∈ s, (p x 0).uncurry0 = f x) -(fderiv_within : ∀ (m : ℕ) (hm : (m : ℕ∞) < n), ∀ x ∈ s, - has_fderiv_within_at (λ y, p y m) (p x m.succ).curry_left s x) -(cont : ∀ (m : ℕ) (hm : (m : ℕ∞) ≤ n), continuous_on (λ x, p x m) s) - -lemma has_ftaylor_series_up_to_on.zero_eq' - (h : has_ftaylor_series_up_to_on n f p s) {x : E} (hx : x ∈ s) : - p x 0 = (continuous_multilinear_curry_fin0 𝕜 E F).symm (f x) := -by { rw ← h.zero_eq x hx, symmetry, exact continuous_multilinear_map.uncurry0_curry0 _ } - -/-- If two functions coincide on a set `s`, then a Taylor series for the first one is as well a -Taylor series for the second one. -/ -lemma has_ftaylor_series_up_to_on.congr - (h : has_ftaylor_series_up_to_on n f p s) (h₁ : ∀ x ∈ s, f₁ x = f x) : - has_ftaylor_series_up_to_on n f₁ p s := -begin - refine ⟨λ x hx, _, h.fderiv_within, h.cont⟩, - rw h₁ x hx, - exact h.zero_eq x hx -end - -lemma has_ftaylor_series_up_to_on.mono - (h : has_ftaylor_series_up_to_on n f p s) {t : set E} (hst : t ⊆ s) : - has_ftaylor_series_up_to_on n f p t := -⟨λ x hx, h.zero_eq x (hst hx), -λ m hm x hx, (h.fderiv_within m hm x (hst hx)).mono hst, -λ m hm, (h.cont m hm).mono hst⟩ - -lemma has_ftaylor_series_up_to_on.of_le - (h : has_ftaylor_series_up_to_on n f p s) (hmn : m ≤ n) : - has_ftaylor_series_up_to_on m f p s := -⟨h.zero_eq, -λ k hk x hx, h.fderiv_within k (lt_of_lt_of_le hk hmn) x hx, -λ k hk, h.cont k (le_trans hk hmn)⟩ - -lemma has_ftaylor_series_up_to_on.continuous_on - (h : has_ftaylor_series_up_to_on n f p s) : continuous_on f s := -begin - have := (h.cont 0 bot_le).congr (λ x hx, (h.zero_eq' hx).symm), - rwa linear_isometry_equiv.comp_continuous_on_iff at this -end - -lemma has_ftaylor_series_up_to_on_zero_iff : - has_ftaylor_series_up_to_on 0 f p s ↔ continuous_on f s ∧ (∀ x ∈ s, (p x 0).uncurry0 = f x) := -begin - refine ⟨λ H, ⟨H.continuous_on, H.zero_eq⟩, - λ H, ⟨H.2, λ m hm, false.elim (not_le.2 hm bot_le), _⟩⟩, - assume m hm, - obtain rfl : m = 0, by exact_mod_cast (hm.antisymm (zero_le _)), - have : ∀ x ∈ s, p x 0 = (continuous_multilinear_curry_fin0 𝕜 E F).symm (f x), - by { assume x hx, rw ← H.2 x hx, symmetry, exact continuous_multilinear_map.uncurry0_curry0 _ }, - rw [continuous_on_congr this, linear_isometry_equiv.comp_continuous_on_iff], - exact H.1 -end - -lemma has_ftaylor_series_up_to_on_top_iff : - (has_ftaylor_series_up_to_on ∞ f p s) ↔ (∀ (n : ℕ), has_ftaylor_series_up_to_on n f p s) := -begin - split, - { assume H n, exact H.of_le le_top }, - { assume H, - split, - { exact (H 0).zero_eq }, - { assume m hm, - apply (H m.succ).fderiv_within m (with_top.coe_lt_coe.2 (lt_add_one m)) }, - { assume m hm, - apply (H m).cont m le_rfl } } -end - -/-- If a function has a Taylor series at order at least `1`, then the term of order `1` of this -series is a derivative of `f`. -/ -lemma has_ftaylor_series_up_to_on.has_fderiv_within_at - (h : has_ftaylor_series_up_to_on n f p s) (hn : 1 ≤ n) (hx : x ∈ s) : - has_fderiv_within_at f (continuous_multilinear_curry_fin1 𝕜 E F (p x 1)) s x := -begin - have A : ∀ y ∈ s, f y = (continuous_multilinear_curry_fin0 𝕜 E F) (p y 0), - { assume y hy, rw ← h.zero_eq y hy, refl }, - suffices H : has_fderiv_within_at - (λ y, continuous_multilinear_curry_fin0 𝕜 E F (p y 0)) - (continuous_multilinear_curry_fin1 𝕜 E F (p x 1)) s x, - by exact H.congr A (A x hx), - rw linear_isometry_equiv.comp_has_fderiv_within_at_iff', - have : ((0 : ℕ) : ℕ∞) < n := - lt_of_lt_of_le (with_top.coe_lt_coe.2 nat.zero_lt_one) hn, - convert h.fderiv_within _ this x hx, - ext y v, - change (p x 1) (snoc 0 y) = (p x 1) (cons y v), - unfold_coes, - congr' with i, - rw unique.eq_default i, - refl -end - -lemma has_ftaylor_series_up_to_on.differentiable_on - (h : has_ftaylor_series_up_to_on n f p s) (hn : 1 ≤ n) : differentiable_on 𝕜 f s := -λ x hx, (h.has_fderiv_within_at hn hx).differentiable_within_at - -/-- If a function has a Taylor series at order at least `1` on a neighborhood of `x`, then the term -of order `1` of this series is a derivative of `f` at `x`. -/ -lemma has_ftaylor_series_up_to_on.has_fderiv_at - (h : has_ftaylor_series_up_to_on n f p s) (hn : 1 ≤ n) (hx : s ∈ 𝓝 x) : - has_fderiv_at f (continuous_multilinear_curry_fin1 𝕜 E F (p x 1)) x := -(h.has_fderiv_within_at hn (mem_of_mem_nhds hx)).has_fderiv_at hx - -/-- If a function has a Taylor series at order at least `1` on a neighborhood of `x`, then -in a neighborhood of `x`, the term of order `1` of this series is a derivative of `f`. -/ -lemma has_ftaylor_series_up_to_on.eventually_has_fderiv_at - (h : has_ftaylor_series_up_to_on n f p s) (hn : 1 ≤ n) (hx : s ∈ 𝓝 x) : - ∀ᶠ y in 𝓝 x, has_fderiv_at f (continuous_multilinear_curry_fin1 𝕜 E F (p y 1)) y := -(eventually_eventually_nhds.2 hx).mono $ λ y hy, h.has_fderiv_at hn hy - -/-- If a function has a Taylor series at order at least `1` on a neighborhood of `x`, then -it is differentiable at `x`. -/ -lemma has_ftaylor_series_up_to_on.differentiable_at - (h : has_ftaylor_series_up_to_on n f p s) (hn : 1 ≤ n) (hx : s ∈ 𝓝 x) : - differentiable_at 𝕜 f x := -(h.has_fderiv_at hn hx).differentiable_at - -/-- `p` is a Taylor series of `f` up to `n+1` if and only if `p` is a Taylor series up to `n`, and -`p (n + 1)` is a derivative of `p n`. -/ -theorem has_ftaylor_series_up_to_on_succ_iff_left {n : ℕ} : - has_ftaylor_series_up_to_on (n + 1) f p s ↔ - has_ftaylor_series_up_to_on n f p s - ∧ (∀ x ∈ s, has_fderiv_within_at (λ y, p y n) (p x n.succ).curry_left s x) - ∧ continuous_on (λ x, p x (n + 1)) s := -begin - split, - { assume h, - exact ⟨h.of_le (with_top.coe_le_coe.2 (nat.le_succ n)), - h.fderiv_within _ (with_top.coe_lt_coe.2 (lt_add_one n)), - h.cont (n + 1) le_rfl⟩ }, - { assume h, - split, - { exact h.1.zero_eq }, - { assume m hm, - by_cases h' : m < n, - { exact h.1.fderiv_within m (with_top.coe_lt_coe.2 h') }, - { have : m = n := nat.eq_of_lt_succ_of_not_lt (with_top.coe_lt_coe.1 hm) h', - rw this, - exact h.2.1 } }, - { assume m hm, - by_cases h' : m ≤ n, - { apply h.1.cont m (with_top.coe_le_coe.2 h') }, - { have : m = (n + 1) := le_antisymm (with_top.coe_le_coe.1 hm) (not_le.1 h'), - rw this, - exact h.2.2 } } } -end - -/-- `p` is a Taylor series of `f` up to `n+1` if and only if `p.shift` is a Taylor series up to `n` -for `p 1`, which is a derivative of `f`. -/ -theorem has_ftaylor_series_up_to_on_succ_iff_right {n : ℕ} : - has_ftaylor_series_up_to_on ((n + 1) : ℕ) f p s ↔ - (∀ x ∈ s, (p x 0).uncurry0 = f x) - ∧ (∀ x ∈ s, has_fderiv_within_at (λ y, p y 0) (p x 1).curry_left s x) - ∧ has_ftaylor_series_up_to_on n - (λ x, continuous_multilinear_curry_fin1 𝕜 E F (p x 1)) (λ x, (p x).shift) s := -begin - split, - { assume H, - refine ⟨H.zero_eq, H.fderiv_within 0 (with_top.coe_lt_coe.2 (nat.succ_pos n)), _⟩, - split, - { assume x hx, refl }, - { assume m (hm : (m : ℕ∞) < n) x (hx : x ∈ s), - have A : (m.succ : ℕ∞) < n.succ, - by { rw with_top.coe_lt_coe at ⊢ hm, exact nat.lt_succ_iff.mpr hm }, - change has_fderiv_within_at - ((continuous_multilinear_curry_right_equiv' 𝕜 m E F).symm - ∘ (λ (y : E), p y m.succ)) - (p x m.succ.succ).curry_right.curry_left s x, - rw linear_isometry_equiv.comp_has_fderiv_within_at_iff', - convert H.fderiv_within _ A x hx, - ext y v, - change (p x m.succ.succ) (snoc (cons y (init v)) (v (last _))) - = (p x (nat.succ (nat.succ m))) (cons y v), - rw [← cons_snoc_eq_snoc_cons, snoc_init_self] }, - { assume m (hm : (m : ℕ∞) ≤ n), - have A : (m.succ : ℕ∞) ≤ n.succ, - by { rw with_top.coe_le_coe at ⊢ hm, exact nat.pred_le_iff.mp hm }, - change continuous_on ((continuous_multilinear_curry_right_equiv' 𝕜 m E F).symm - ∘ (λ (y : E), p y m.succ)) s, - rw linear_isometry_equiv.comp_continuous_on_iff, - exact H.cont _ A } }, - { rintros ⟨Hzero_eq, Hfderiv_zero, Htaylor⟩, - split, - { exact Hzero_eq }, - { assume m (hm : (m : ℕ∞) < n.succ) x (hx : x ∈ s), - cases m, - { exact Hfderiv_zero x hx }, - { have A : (m : ℕ∞) < n, - by { rw with_top.coe_lt_coe at hm ⊢, exact nat.lt_of_succ_lt_succ hm }, - have : has_fderiv_within_at ((continuous_multilinear_curry_right_equiv' 𝕜 m E F).symm - ∘ (λ (y : E), p y m.succ)) ((p x).shift m.succ).curry_left s x := - Htaylor.fderiv_within _ A x hx, - rw linear_isometry_equiv.comp_has_fderiv_within_at_iff' at this, - convert this, - ext y v, - change (p x (nat.succ (nat.succ m))) (cons y v) - = (p x m.succ.succ) (snoc (cons y (init v)) (v (last _))), - rw [← cons_snoc_eq_snoc_cons, snoc_init_self] } }, - { assume m (hm : (m : ℕ∞) ≤ n.succ), - cases m, - { have : differentiable_on 𝕜 (λ x, p x 0) s := - λ x hx, (Hfderiv_zero x hx).differentiable_within_at, - exact this.continuous_on }, - { have A : (m : ℕ∞) ≤ n, - by { rw with_top.coe_le_coe at hm ⊢, exact nat.lt_succ_iff.mp hm }, - have : continuous_on ((continuous_multilinear_curry_right_equiv' 𝕜 m E F).symm - ∘ (λ (y : E), p y m.succ)) s := - Htaylor.cont _ A, - rwa linear_isometry_equiv.comp_continuous_on_iff at this } } } -end - -/-! ### Smooth functions within a set around a point -/ - -variable (𝕜) - -/-- A function is continuously differentiable up to order `n` within a set `s` at a point `x` if -it admits continuous derivatives up to order `n` in a neighborhood of `x` in `s ∪ {x}`. -For `n = ∞`, we only require that this holds up to any finite order (where the neighborhood may -depend on the finite order we consider). - -For instance, a real function which is `C^m` on `(-1/m, 1/m)` for each natural `m`, but not -better, is `C^∞` at `0` within `univ`. --/ -def cont_diff_within_at (n : ℕ∞) (f : E → F) (s : set E) (x : E) : Prop := -∀ (m : ℕ), (m : ℕ∞) ≤ n → - ∃ u ∈ 𝓝[insert x s] x, ∃ p : E → formal_multilinear_series 𝕜 E F, - has_ftaylor_series_up_to_on m f p u - -variable {𝕜} - -lemma cont_diff_within_at_nat {n : ℕ} : - cont_diff_within_at 𝕜 n f s x ↔ - ∃ u ∈ 𝓝[insert x s] x, ∃ p : E → formal_multilinear_series 𝕜 E F, - has_ftaylor_series_up_to_on n f p u := -⟨λ H, H n le_rfl, λ ⟨u, hu, p, hp⟩ m hm, ⟨u, hu, p, hp.of_le hm⟩⟩ - -lemma cont_diff_within_at.of_le - (h : cont_diff_within_at 𝕜 n f s x) (hmn : m ≤ n) : - cont_diff_within_at 𝕜 m f s x := -λ k hk, h k (le_trans hk hmn) - -lemma cont_diff_within_at_iff_forall_nat_le : - cont_diff_within_at 𝕜 n f s x ↔ ∀ m : ℕ, ↑m ≤ n → cont_diff_within_at 𝕜 m f s x := -⟨λ H m hm, H.of_le hm, λ H m hm, H m hm _ le_rfl⟩ - -lemma cont_diff_within_at_top : - cont_diff_within_at 𝕜 ∞ f s x ↔ ∀ (n : ℕ), cont_diff_within_at 𝕜 n f s x := -cont_diff_within_at_iff_forall_nat_le.trans $ by simp only [forall_prop_of_true, le_top] - -lemma cont_diff_within_at.continuous_within_at - (h : cont_diff_within_at 𝕜 n f s x) : continuous_within_at f s x := -begin - rcases h 0 bot_le with ⟨u, hu, p, H⟩, - rw [mem_nhds_within_insert] at hu, - exact (H.continuous_on.continuous_within_at hu.1).mono_of_mem hu.2 -end - -lemma cont_diff_within_at.congr_of_eventually_eq - (h : cont_diff_within_at 𝕜 n f s x) (h₁ : f₁ =ᶠ[𝓝[s] x] f) (hx : f₁ x = f x) : - cont_diff_within_at 𝕜 n f₁ s x := -λ m hm, let ⟨u, hu, p, H⟩ := h m hm in -⟨{x ∈ u | f₁ x = f x}, filter.inter_mem hu (mem_nhds_within_insert.2 ⟨hx, h₁⟩), p, - (H.mono (sep_subset _ _)).congr (λ _, and.right)⟩ - -lemma cont_diff_within_at.congr_of_eventually_eq_insert - (h : cont_diff_within_at 𝕜 n f s x) (h₁ : f₁ =ᶠ[𝓝[insert x s] x] f) : - cont_diff_within_at 𝕜 n f₁ s x := -h.congr_of_eventually_eq (nhds_within_mono x (subset_insert x s) h₁) - (mem_of_mem_nhds_within (mem_insert x s) h₁ : _) - -lemma cont_diff_within_at.congr_of_eventually_eq' - (h : cont_diff_within_at 𝕜 n f s x) (h₁ : f₁ =ᶠ[𝓝[s] x] f) (hx : x ∈ s) : - cont_diff_within_at 𝕜 n f₁ s x := -h.congr_of_eventually_eq h₁ $ h₁.self_of_nhds_within hx - -lemma filter.eventually_eq.cont_diff_within_at_iff - (h₁ : f₁ =ᶠ[𝓝[s] x] f) (hx : f₁ x = f x) : - cont_diff_within_at 𝕜 n f₁ s x ↔ cont_diff_within_at 𝕜 n f s x := -⟨λ H, cont_diff_within_at.congr_of_eventually_eq H h₁.symm hx.symm, -λ H, H.congr_of_eventually_eq h₁ hx⟩ - -lemma cont_diff_within_at.congr - (h : cont_diff_within_at 𝕜 n f s x) (h₁ : ∀ y ∈ s, f₁ y = f y) (hx : f₁ x = f x) : - cont_diff_within_at 𝕜 n f₁ s x := -h.congr_of_eventually_eq (filter.eventually_eq_of_mem self_mem_nhds_within h₁) hx - -lemma cont_diff_within_at.congr' - (h : cont_diff_within_at 𝕜 n f s x) (h₁ : ∀ y ∈ s, f₁ y = f y) (hx : x ∈ s) : - cont_diff_within_at 𝕜 n f₁ s x := -h.congr h₁ (h₁ _ hx) - -lemma cont_diff_within_at.mono_of_mem - (h : cont_diff_within_at 𝕜 n f s x) {t : set E} (hst : s ∈ 𝓝[t] x) : - cont_diff_within_at 𝕜 n f t x := -begin - assume m hm, - rcases h m hm with ⟨u, hu, p, H⟩, - exact ⟨u, nhds_within_le_of_mem (insert_mem_nhds_within_insert hst) hu, p, H⟩ -end - -lemma cont_diff_within_at.mono - (h : cont_diff_within_at 𝕜 n f s x) {t : set E} (hst : t ⊆ s) : - cont_diff_within_at 𝕜 n f t x := -h.mono_of_mem $ filter.mem_of_superset self_mem_nhds_within hst - -lemma cont_diff_within_at.congr_nhds - (h : cont_diff_within_at 𝕜 n f s x) {t : set E} (hst : 𝓝[s] x = 𝓝[t] x) : - cont_diff_within_at 𝕜 n f t x := -h.mono_of_mem $ hst ▸ self_mem_nhds_within - -lemma cont_diff_within_at_congr_nhds {t : set E} (hst : 𝓝[s] x = 𝓝[t] x) : - cont_diff_within_at 𝕜 n f s x ↔ cont_diff_within_at 𝕜 n f t x := -⟨λ h, h.congr_nhds hst, λ h, h.congr_nhds hst.symm⟩ - -lemma cont_diff_within_at_inter' (h : t ∈ 𝓝[s] x) : - cont_diff_within_at 𝕜 n f (s ∩ t) x ↔ cont_diff_within_at 𝕜 n f s x := -cont_diff_within_at_congr_nhds $ eq.symm $ nhds_within_restrict'' _ h - -lemma cont_diff_within_at_inter (h : t ∈ 𝓝 x) : - cont_diff_within_at 𝕜 n f (s ∩ t) x ↔ cont_diff_within_at 𝕜 n f s x := -cont_diff_within_at_inter' (mem_nhds_within_of_mem_nhds h) - -lemma cont_diff_within_at_insert {y : E} : - cont_diff_within_at 𝕜 n f (insert y s) x ↔ cont_diff_within_at 𝕜 n f s x := -begin - simp_rw [cont_diff_within_at], - rcases eq_or_ne x y with rfl|h, - { simp_rw [insert_eq_of_mem (mem_insert _ _)] }, - simp_rw [insert_comm x y, nhds_within_insert_of_ne h] -end - -alias cont_diff_within_at_insert ↔ cont_diff_within_at.of_insert cont_diff_within_at.insert' - -lemma cont_diff_within_at.insert (h : cont_diff_within_at 𝕜 n f s x) : - cont_diff_within_at 𝕜 n f (insert x s) x := -h.insert' - -/-- If a function is `C^n` within a set at a point, with `n ≥ 1`, then it is differentiable -within this set at this point. -/ -lemma cont_diff_within_at.differentiable_within_at' - (h : cont_diff_within_at 𝕜 n f s x) (hn : 1 ≤ n) : - differentiable_within_at 𝕜 f (insert x s) x := -begin - rcases h 1 hn with ⟨u, hu, p, H⟩, - rcases mem_nhds_within.1 hu with ⟨t, t_open, xt, tu⟩, - rw inter_comm at tu, - have := ((H.mono tu).differentiable_on le_rfl) x ⟨mem_insert x s, xt⟩, - exact (differentiable_within_at_inter (is_open.mem_nhds t_open xt)).1 this, -end - -lemma cont_diff_within_at.differentiable_within_at - (h : cont_diff_within_at 𝕜 n f s x) (hn : 1 ≤ n) : - differentiable_within_at 𝕜 f s x := -(h.differentiable_within_at' hn).mono (subset_insert x s) - -/-- A function is `C^(n + 1)` on a domain iff locally, it has a derivative which is `C^n`. -/ -theorem cont_diff_within_at_succ_iff_has_fderiv_within_at {n : ℕ} : - cont_diff_within_at 𝕜 ((n + 1) : ℕ) f s x - ↔ ∃ u ∈ 𝓝[insert x s] x, ∃ f' : E → (E →L[𝕜] F), - (∀ x ∈ u, has_fderiv_within_at f (f' x) u x) ∧ (cont_diff_within_at 𝕜 n f' u x) := -begin - split, - { assume h, - rcases h n.succ le_rfl with ⟨u, hu, p, Hp⟩, - refine ⟨u, hu, λ y, (continuous_multilinear_curry_fin1 𝕜 E F) (p y 1), - λ y hy, Hp.has_fderiv_within_at (with_top.coe_le_coe.2 (nat.le_add_left 1 n)) hy, _⟩, - assume m hm, - refine ⟨u, _, λ (y : E), (p y).shift, _⟩, - { convert self_mem_nhds_within, - have : x ∈ insert x s, by simp, - exact (insert_eq_of_mem (mem_of_mem_nhds_within this hu)) }, - { rw has_ftaylor_series_up_to_on_succ_iff_right at Hp, - exact Hp.2.2.of_le hm } }, - { rintros ⟨u, hu, f', f'_eq_deriv, Hf'⟩, - rw cont_diff_within_at_nat, - rcases Hf' n le_rfl with ⟨v, hv, p', Hp'⟩, - refine ⟨v ∩ u, _, λ x, (p' x).unshift (f x), _⟩, - { apply filter.inter_mem _ hu, - apply nhds_within_le_of_mem hu, - exact nhds_within_mono _ (subset_insert x u) hv }, - { rw has_ftaylor_series_up_to_on_succ_iff_right, - refine ⟨λ y hy, rfl, λ y hy, _, _⟩, - { change has_fderiv_within_at (λ z, (continuous_multilinear_curry_fin0 𝕜 E F).symm (f z)) - ((formal_multilinear_series.unshift (p' y) (f y) 1).curry_left) (v ∩ u) y, - rw linear_isometry_equiv.comp_has_fderiv_within_at_iff', - convert (f'_eq_deriv y hy.2).mono (inter_subset_right v u), - rw ← Hp'.zero_eq y hy.1, - ext z, - change ((p' y 0) (init (@cons 0 (λ i, E) z 0))) (@cons 0 (λ i, E) z 0 (last 0)) - = ((p' y 0) 0) z, - unfold_coes, - congr }, - { convert (Hp'.mono (inter_subset_left v u)).congr (λ x hx, Hp'.zero_eq x hx.1), - { ext x y, - change p' x 0 (init (@snoc 0 (λ i : fin 1, E) 0 y)) y = p' x 0 0 y, - rw init_snoc }, - { ext x k v y, - change p' x k (init (@snoc k (λ i : fin k.succ, E) v y)) - (@snoc k (λ i : fin k.succ, E) v y (last k)) = p' x k v y, - rw [snoc_last, init_snoc] } } } } -end - -/-- A version of `cont_diff_within_at_succ_iff_has_fderiv_within_at` where all derivatives - are taken within the same set. -/ -lemma cont_diff_within_at_succ_iff_has_fderiv_within_at' {n : ℕ} : - cont_diff_within_at 𝕜 (n + 1 : ℕ) f s x - ↔ ∃ u ∈ 𝓝[insert x s] x, u ⊆ insert x s ∧ ∃ f' : E → E →L[𝕜] F, - (∀ x ∈ u, has_fderiv_within_at f (f' x) s x) ∧ cont_diff_within_at 𝕜 n f' s x := -begin - refine ⟨λ hf, _, _⟩, - { obtain ⟨u, hu, f', huf', hf'⟩ := cont_diff_within_at_succ_iff_has_fderiv_within_at.mp hf, - obtain ⟨w, hw, hxw, hwu⟩ := mem_nhds_within.mp hu, - rw [inter_comm] at hwu, - refine ⟨insert x s ∩ w, inter_mem_nhds_within _ (hw.mem_nhds hxw), inter_subset_left _ _, - f', λ y hy, _, _⟩, - { refine ((huf' y $ hwu hy).mono hwu).mono_of_mem _, - refine mem_of_superset _ (inter_subset_inter_left _ (subset_insert _ _)), - refine inter_mem_nhds_within _ (hw.mem_nhds hy.2) }, - { exact hf'.mono_of_mem (nhds_within_mono _ (subset_insert _ _) hu) } }, - { rw [← cont_diff_within_at_insert, cont_diff_within_at_succ_iff_has_fderiv_within_at, - insert_eq_of_mem (mem_insert _ _)], - rintro ⟨u, hu, hus, f', huf', hf'⟩, - refine ⟨u, hu, f', λ y hy, (huf' y hy).insert'.mono hus, hf'.insert.mono hus⟩ } -end - -/-! ### Smooth functions within a set -/ - -variable (𝕜) - -/-- A function is continuously differentiable up to `n` on `s` if, for any point `x` in `s`, it -admits continuous derivatives up to order `n` on a neighborhood of `x` in `s`. - -For `n = ∞`, we only require that this holds up to any finite order (where the neighborhood may -depend on the finite order we consider). --/ -def cont_diff_on (n : ℕ∞) (f : E → F) (s : set E) : Prop := -∀ x ∈ s, cont_diff_within_at 𝕜 n f s x - -variable {𝕜} - -lemma cont_diff_on.cont_diff_within_at (h : cont_diff_on 𝕜 n f s) (hx : x ∈ s) : - cont_diff_within_at 𝕜 n f s x := -h x hx - -lemma cont_diff_within_at.cont_diff_on {m : ℕ} - (hm : (m : ℕ∞) ≤ n) (h : cont_diff_within_at 𝕜 n f s x) : - ∃ u ∈ 𝓝[insert x s] x, u ⊆ insert x s ∧ cont_diff_on 𝕜 m f u := -begin - rcases h m hm with ⟨u, u_nhd, p, hp⟩, - refine ⟨u ∩ insert x s, filter.inter_mem u_nhd self_mem_nhds_within, - inter_subset_right _ _, _⟩, - assume y hy m' hm', - refine ⟨u ∩ insert x s, _, p, (hp.mono (inter_subset_left _ _)).of_le hm'⟩, - convert self_mem_nhds_within, - exact insert_eq_of_mem hy -end - -protected lemma cont_diff_within_at.eventually {n : ℕ} - (h : cont_diff_within_at 𝕜 n f s x) : - ∀ᶠ y in 𝓝[insert x s] x, cont_diff_within_at 𝕜 n f s y := -begin - rcases h.cont_diff_on le_rfl with ⟨u, hu, hu_sub, hd⟩, - have : ∀ᶠ (y : E) in 𝓝[insert x s] x, u ∈ 𝓝[insert x s] y ∧ y ∈ u, - from (eventually_nhds_within_nhds_within.2 hu).and hu, - refine this.mono (λ y hy, (hd y hy.2).mono_of_mem _), - exact nhds_within_mono y (subset_insert _ _) hy.1 -end - -lemma cont_diff_on.of_le (h : cont_diff_on 𝕜 n f s) (hmn : m ≤ n) : - cont_diff_on 𝕜 m f s := -λ x hx, (h x hx).of_le hmn - -lemma cont_diff_on.of_succ {n : ℕ} (h : cont_diff_on 𝕜 (n + 1) f s) : cont_diff_on 𝕜 n f s := -h.of_le $ with_top.coe_le_coe.mpr le_self_add - -lemma cont_diff_on.one_of_succ {n : ℕ} (h : cont_diff_on 𝕜 (n + 1) f s) : cont_diff_on 𝕜 1 f s := -h.of_le $ with_top.coe_le_coe.mpr le_add_self - -lemma cont_diff_on_iff_forall_nat_le : - cont_diff_on 𝕜 n f s ↔ ∀ m : ℕ, ↑m ≤ n → cont_diff_on 𝕜 m f s := -⟨λ H m hm, H.of_le hm, λ H x hx m hm, H m hm x hx m le_rfl⟩ - -lemma cont_diff_on_top : - cont_diff_on 𝕜 ∞ f s ↔ ∀ (n : ℕ), cont_diff_on 𝕜 n f s := -cont_diff_on_iff_forall_nat_le.trans $ by simp only [le_top, forall_prop_of_true] - -lemma cont_diff_on_all_iff_nat : - (∀ n, cont_diff_on 𝕜 n f s) ↔ (∀ n : ℕ, cont_diff_on 𝕜 n f s) := -begin - refine ⟨λ H n, H n, _⟩, - rintro H (_|n), - exacts [cont_diff_on_top.2 H, H n] -end - -lemma cont_diff_on.continuous_on - (h : cont_diff_on 𝕜 n f s) : continuous_on f s := -λ x hx, (h x hx).continuous_within_at - -lemma cont_diff_on.congr - (h : cont_diff_on 𝕜 n f s) (h₁ : ∀ x ∈ s, f₁ x = f x) : - cont_diff_on 𝕜 n f₁ s := -λ x hx, (h x hx).congr h₁ (h₁ x hx) - -lemma cont_diff_on_congr (h₁ : ∀ x ∈ s, f₁ x = f x) : - cont_diff_on 𝕜 n f₁ s ↔ cont_diff_on 𝕜 n f s := -⟨λ H, H.congr (λ x hx, (h₁ x hx).symm), λ H, H.congr h₁⟩ - -lemma cont_diff_on.mono - (h : cont_diff_on 𝕜 n f s) {t : set E} (hst : t ⊆ s) : - cont_diff_on 𝕜 n f t := -λ x hx, (h x (hst hx)).mono hst - -lemma cont_diff_on.congr_mono - (hf : cont_diff_on 𝕜 n f s) (h₁ : ∀ x ∈ s₁, f₁ x = f x) (hs : s₁ ⊆ s) : - cont_diff_on 𝕜 n f₁ s₁ := -(hf.mono hs).congr h₁ - -/-- If a function is `C^n` on a set with `n ≥ 1`, then it is differentiable there. -/ -lemma cont_diff_on.differentiable_on - (h : cont_diff_on 𝕜 n f s) (hn : 1 ≤ n) : differentiable_on 𝕜 f s := -λ x hx, (h x hx).differentiable_within_at hn - -/-- If a function is `C^n` around each point in a set, then it is `C^n` on the set. -/ -lemma cont_diff_on_of_locally_cont_diff_on - (h : ∀ x ∈ s, ∃u, is_open u ∧ x ∈ u ∧ cont_diff_on 𝕜 n f (s ∩ u)) : - cont_diff_on 𝕜 n f s := -begin - assume x xs, - rcases h x xs with ⟨u, u_open, xu, hu⟩, - apply (cont_diff_within_at_inter _).1 (hu x ⟨xs, xu⟩), - exact is_open.mem_nhds u_open xu -end - -/-- A function is `C^(n + 1)` on a domain iff locally, it has a derivative which is `C^n`. -/ -theorem cont_diff_on_succ_iff_has_fderiv_within_at {n : ℕ} : - cont_diff_on 𝕜 ((n + 1) : ℕ) f s - ↔ ∀ x ∈ s, ∃ u ∈ 𝓝[insert x s] x, ∃ f' : E → (E →L[𝕜] F), - (∀ x ∈ u, has_fderiv_within_at f (f' x) u x) ∧ (cont_diff_on 𝕜 n f' u) := -begin - split, - { assume h x hx, - rcases (h x hx) n.succ le_rfl with ⟨u, hu, p, Hp⟩, - refine ⟨u, hu, λ y, (continuous_multilinear_curry_fin1 𝕜 E F) (p y 1), - λ y hy, Hp.has_fderiv_within_at (with_top.coe_le_coe.2 (nat.le_add_left 1 n)) hy, _⟩, - rw has_ftaylor_series_up_to_on_succ_iff_right at Hp, - assume z hz m hm, - refine ⟨u, _, λ (x : E), (p x).shift, Hp.2.2.of_le hm⟩, - convert self_mem_nhds_within, - exact insert_eq_of_mem hz, }, - { assume h x hx, - rw cont_diff_within_at_succ_iff_has_fderiv_within_at, - rcases h x hx with ⟨u, u_nhbd, f', hu, hf'⟩, - have : x ∈ u := mem_of_mem_nhds_within (mem_insert _ _) u_nhbd, - exact ⟨u, u_nhbd, f', hu, hf' x this⟩ } -end - -/-! ### Iterated derivative within a set -/ -variable (𝕜) - -/-- -The `n`-th derivative of a function along a set, defined inductively by saying that the `n+1`-th -derivative of `f` is the derivative of the `n`-th derivative of `f` along this set, together with -an uncurrying step to see it as a multilinear map in `n+1` variables.. --/ -noncomputable def iterated_fderiv_within (n : ℕ) (f : E → F) (s : set E) : - E → (E [×n]→L[𝕜] F) := -nat.rec_on n - (λ x, continuous_multilinear_map.curry0 𝕜 E (f x)) - (λ n rec x, continuous_linear_map.uncurry_left (fderiv_within 𝕜 rec s x)) - -/-- Formal Taylor series associated to a function within a set. -/ -def ftaylor_series_within (f : E → F) (s : set E) (x : E) : formal_multilinear_series 𝕜 E F := -λ n, iterated_fderiv_within 𝕜 n f s x - -variable {𝕜} - -@[simp] lemma iterated_fderiv_within_zero_apply (m : (fin 0) → E) : - (iterated_fderiv_within 𝕜 0 f s x : ((fin 0) → E) → F) m = f x := rfl - -lemma iterated_fderiv_within_zero_eq_comp : - iterated_fderiv_within 𝕜 0 f s = (continuous_multilinear_curry_fin0 𝕜 E F).symm ∘ f := rfl - -@[simp] lemma norm_iterated_fderiv_within_zero : - ‖iterated_fderiv_within 𝕜 0 f s x‖ = ‖f x‖ := -by rw [iterated_fderiv_within_zero_eq_comp, linear_isometry_equiv.norm_map] - -lemma iterated_fderiv_within_succ_apply_left {n : ℕ} (m : fin (n + 1) → E): - (iterated_fderiv_within 𝕜 (n + 1) f s x : (fin (n + 1) → E) → F) m - = (fderiv_within 𝕜 (iterated_fderiv_within 𝕜 n f s) s x : E → (E [×n]→L[𝕜] F)) - (m 0) (tail m) := rfl - -/-- Writing explicitly the `n+1`-th derivative as the composition of a currying linear equiv, -and the derivative of the `n`-th derivative. -/ -lemma iterated_fderiv_within_succ_eq_comp_left {n : ℕ} : - iterated_fderiv_within 𝕜 (n + 1) f s = - (continuous_multilinear_curry_left_equiv 𝕜 (λ(i : fin (n + 1)), E) F) - ∘ (fderiv_within 𝕜 (iterated_fderiv_within 𝕜 n f s) s) := rfl - -lemma norm_fderiv_within_iterated_fderiv_within {n : ℕ} : - ‖fderiv_within 𝕜 (iterated_fderiv_within 𝕜 n f s) s x‖ = - ‖iterated_fderiv_within 𝕜 (n + 1) f s x‖ := -by rw [iterated_fderiv_within_succ_eq_comp_left, linear_isometry_equiv.norm_map] - -theorem iterated_fderiv_within_succ_apply_right {n : ℕ} - (hs : unique_diff_on 𝕜 s) (hx : x ∈ s) (m : fin (n + 1) → E) : - (iterated_fderiv_within 𝕜 (n + 1) f s x : (fin (n + 1) → E) → F) m - = iterated_fderiv_within 𝕜 n (λy, fderiv_within 𝕜 f s y) s x (init m) (m (last n)) := -begin - induction n with n IH generalizing x, - { rw [iterated_fderiv_within_succ_eq_comp_left, iterated_fderiv_within_zero_eq_comp, - iterated_fderiv_within_zero_apply, - function.comp_apply, linear_isometry_equiv.comp_fderiv_within _ (hs x hx)], - refl }, - { let I := continuous_multilinear_curry_right_equiv' 𝕜 n E F, - have A : ∀ y ∈ s, iterated_fderiv_within 𝕜 n.succ f s y - = (I ∘ (iterated_fderiv_within 𝕜 n (λy, fderiv_within 𝕜 f s y) s)) y, - by { assume y hy, ext m, rw @IH m y hy, refl }, - calc - (iterated_fderiv_within 𝕜 (n+2) f s x : (fin (n+2) → E) → F) m = - (fderiv_within 𝕜 (iterated_fderiv_within 𝕜 n.succ f s) s x - : E → (E [×(n + 1)]→L[𝕜] F)) (m 0) (tail m) : rfl - ... = (fderiv_within 𝕜 (I ∘ (iterated_fderiv_within 𝕜 n (fderiv_within 𝕜 f s) s)) s x - : E → (E [×(n + 1)]→L[𝕜] F)) (m 0) (tail m) : - by rw fderiv_within_congr (hs x hx) A (A x hx) - ... = (I ∘ fderiv_within 𝕜 ((iterated_fderiv_within 𝕜 n (fderiv_within 𝕜 f s) s)) s x - : E → (E [×(n + 1)]→L[𝕜] F)) (m 0) (tail m) : - by { rw linear_isometry_equiv.comp_fderiv_within _ (hs x hx), refl } - ... = (fderiv_within 𝕜 ((iterated_fderiv_within 𝕜 n (λ y, fderiv_within 𝕜 f s y) s)) s x - : E → (E [×n]→L[𝕜] (E →L[𝕜] F))) (m 0) (init (tail m)) ((tail m) (last n)) : rfl - ... = iterated_fderiv_within 𝕜 (nat.succ n) (λ y, fderiv_within 𝕜 f s y) s x - (init m) (m (last (n + 1))) : - by { rw [iterated_fderiv_within_succ_apply_left, tail_init_eq_init_tail], refl } } -end - -/-- Writing explicitly the `n+1`-th derivative as the composition of a currying linear equiv, -and the `n`-th derivative of the derivative. -/ -lemma iterated_fderiv_within_succ_eq_comp_right {n : ℕ} (hs : unique_diff_on 𝕜 s) (hx : x ∈ s) : - iterated_fderiv_within 𝕜 (n + 1) f s x = - ((continuous_multilinear_curry_right_equiv' 𝕜 n E F) - ∘ (iterated_fderiv_within 𝕜 n (λy, fderiv_within 𝕜 f s y) s)) x := -by { ext m, rw iterated_fderiv_within_succ_apply_right hs hx, refl } - -lemma norm_iterated_fderiv_within_fderiv_within {n : ℕ} (hs : unique_diff_on 𝕜 s) (hx : x ∈ s) : - ‖iterated_fderiv_within 𝕜 n (fderiv_within 𝕜 f s) s x‖ = - ‖iterated_fderiv_within 𝕜 (n + 1) f s x‖ := -by rw [iterated_fderiv_within_succ_eq_comp_right hs hx, linear_isometry_equiv.norm_map] - -@[simp] lemma iterated_fderiv_within_one_apply - (hs : unique_diff_on 𝕜 s) (hx : x ∈ s) (m : (fin 1) → E) : - (iterated_fderiv_within 𝕜 1 f s x : ((fin 1) → E) → F) m - = (fderiv_within 𝕜 f s x : E → F) (m 0) := -by { rw [iterated_fderiv_within_succ_apply_right hs hx, iterated_fderiv_within_zero_apply], refl } - -/-- If two functions coincide on a set `s` of unique differentiability, then their iterated -differentials within this set coincide. -/ -lemma iterated_fderiv_within_congr {n : ℕ} - (hs : unique_diff_on 𝕜 s) (hL : ∀y∈s, f₁ y = f y) (hx : x ∈ s) : - iterated_fderiv_within 𝕜 n f₁ s x = iterated_fderiv_within 𝕜 n f s x := -begin - induction n with n IH generalizing x, - { ext m, simp [hL x hx] }, - { have : fderiv_within 𝕜 (λ y, iterated_fderiv_within 𝕜 n f₁ s y) s x - = fderiv_within 𝕜 (λ y, iterated_fderiv_within 𝕜 n f s y) s x := - fderiv_within_congr (hs x hx) (λ y hy, IH hy) (IH hx), - ext m, - rw [iterated_fderiv_within_succ_apply_left, iterated_fderiv_within_succ_apply_left, this] } -end - -/-- The iterated differential within a set `s` at a point `x` is not modified if one intersects -`s` with an open set containing `x`. -/ -lemma iterated_fderiv_within_inter_open {n : ℕ} (hu : is_open u) - (hs : unique_diff_on 𝕜 (s ∩ u)) (hx : x ∈ s ∩ u) : - iterated_fderiv_within 𝕜 n f (s ∩ u) x = iterated_fderiv_within 𝕜 n f s x := -begin - induction n with n IH generalizing x, - { ext m, simp }, - { have A : fderiv_within 𝕜 (λ y, iterated_fderiv_within 𝕜 n f (s ∩ u) y) (s ∩ u) x - = fderiv_within 𝕜 (λ y, iterated_fderiv_within 𝕜 n f s y) (s ∩ u) x := - fderiv_within_congr (hs x hx) (λ y hy, IH hy) (IH hx), - have B : fderiv_within 𝕜 (λ y, iterated_fderiv_within 𝕜 n f s y) (s ∩ u) x - = fderiv_within 𝕜 (λ y, iterated_fderiv_within 𝕜 n f s y) s x := - fderiv_within_inter (is_open.mem_nhds hu hx.2) - ((unique_diff_within_at_inter (is_open.mem_nhds hu hx.2)).1 (hs x hx)), - ext m, - rw [iterated_fderiv_within_succ_apply_left, iterated_fderiv_within_succ_apply_left, A, B] } -end - -/-- The iterated differential within a set `s` at a point `x` is not modified if one intersects -`s` with a neighborhood of `x` within `s`. -/ -lemma iterated_fderiv_within_inter' {n : ℕ} - (hu : u ∈ 𝓝[s] x) (hs : unique_diff_on 𝕜 s) (xs : x ∈ s) : - iterated_fderiv_within 𝕜 n f (s ∩ u) x = iterated_fderiv_within 𝕜 n f s x := -begin - obtain ⟨v, v_open, xv, vu⟩ : ∃ v, is_open v ∧ x ∈ v ∧ v ∩ s ⊆ u := mem_nhds_within.1 hu, - have A : (s ∩ u) ∩ v = s ∩ v, - { apply subset.antisymm (inter_subset_inter (inter_subset_left _ _) (subset.refl _)), - exact λ y ⟨ys, yv⟩, ⟨⟨ys, vu ⟨yv, ys⟩⟩, yv⟩ }, - have : iterated_fderiv_within 𝕜 n f (s ∩ v) x = iterated_fderiv_within 𝕜 n f s x := - iterated_fderiv_within_inter_open v_open (hs.inter v_open) ⟨xs, xv⟩, - rw ← this, - have : iterated_fderiv_within 𝕜 n f ((s ∩ u) ∩ v) x = iterated_fderiv_within 𝕜 n f (s ∩ u) x, - { refine iterated_fderiv_within_inter_open v_open _ ⟨⟨xs, vu ⟨xv, xs⟩⟩, xv⟩, - rw A, - exact hs.inter v_open }, - rw A at this, - rw ← this -end - -/-- The iterated differential within a set `s` at a point `x` is not modified if one intersects -`s` with a neighborhood of `x`. -/ -lemma iterated_fderiv_within_inter {n : ℕ} - (hu : u ∈ 𝓝 x) (hs : unique_diff_on 𝕜 s) (xs : x ∈ s) : - iterated_fderiv_within 𝕜 n f (s ∩ u) x = iterated_fderiv_within 𝕜 n f s x := -iterated_fderiv_within_inter' (mem_nhds_within_of_mem_nhds hu) hs xs - -@[simp] lemma cont_diff_on_zero : - cont_diff_on 𝕜 0 f s ↔ continuous_on f s := -begin - refine ⟨λ H, H.continuous_on, λ H, _⟩, - assume x hx m hm, - have : (m : ℕ∞) = 0 := le_antisymm hm bot_le, - rw this, - refine ⟨insert x s, self_mem_nhds_within, ftaylor_series_within 𝕜 f s, _⟩, - rw has_ftaylor_series_up_to_on_zero_iff, - exact ⟨by rwa insert_eq_of_mem hx, λ x hx, by simp [ftaylor_series_within]⟩ -end - -lemma cont_diff_within_at_zero (hx : x ∈ s) : - cont_diff_within_at 𝕜 0 f s x ↔ ∃ u ∈ 𝓝[s] x, continuous_on f (s ∩ u) := -begin - split, - { intros h, - obtain ⟨u, H, p, hp⟩ := h 0 (by norm_num), - refine ⟨u, _, _⟩, - { simpa [hx] using H }, - { simp only [with_top.coe_zero, has_ftaylor_series_up_to_on_zero_iff] at hp, - exact hp.1.mono (inter_subset_right s u) } }, - { rintros ⟨u, H, hu⟩, - rw ← cont_diff_within_at_inter' H, - have h' : x ∈ s ∩ u := ⟨hx, mem_of_mem_nhds_within hx H⟩, - exact (cont_diff_on_zero.mpr hu).cont_diff_within_at h' } -end - -/-- On a set with unique differentiability, any choice of iterated differential has to coincide -with the one we have chosen in `iterated_fderiv_within 𝕜 m f s`. -/ -theorem has_ftaylor_series_up_to_on.eq_ftaylor_series_of_unique_diff_on - (h : has_ftaylor_series_up_to_on n f p s) - {m : ℕ} (hmn : (m : ℕ∞) ≤ n) (hs : unique_diff_on 𝕜 s) (hx : x ∈ s) : - p x m = iterated_fderiv_within 𝕜 m f s x := -begin - induction m with m IH generalizing x, - { rw [h.zero_eq' hx, iterated_fderiv_within_zero_eq_comp] }, - { have A : (m : ℕ∞) < n := lt_of_lt_of_le (with_top.coe_lt_coe.2 (lt_add_one m)) hmn, - have : has_fderiv_within_at (λ (y : E), iterated_fderiv_within 𝕜 m f s y) - (continuous_multilinear_map.curry_left (p x (nat.succ m))) s x := - (h.fderiv_within m A x hx).congr (λ y hy, (IH (le_of_lt A) hy).symm) (IH (le_of_lt A) hx).symm, - rw [iterated_fderiv_within_succ_eq_comp_left, function.comp_apply, - this.fderiv_within (hs x hx)], - exact (continuous_multilinear_map.uncurry_curry_left _).symm } -end - -/-- When a function is `C^n` in a set `s` of unique differentiability, it admits -`ftaylor_series_within 𝕜 f s` as a Taylor series up to order `n` in `s`. -/ -theorem cont_diff_on.ftaylor_series_within - (h : cont_diff_on 𝕜 n f s) (hs : unique_diff_on 𝕜 s) : - has_ftaylor_series_up_to_on n f (ftaylor_series_within 𝕜 f s) s := -begin - split, - { assume x hx, - simp only [ftaylor_series_within, continuous_multilinear_map.uncurry0_apply, - iterated_fderiv_within_zero_apply] }, - { assume m hm x hx, - rcases (h x hx) m.succ (enat.add_one_le_of_lt hm) with ⟨u, hu, p, Hp⟩, - rw insert_eq_of_mem hx at hu, - rcases mem_nhds_within.1 hu with ⟨o, o_open, xo, ho⟩, - rw inter_comm at ho, - have : p x m.succ = ftaylor_series_within 𝕜 f s x m.succ, - { change p x m.succ = iterated_fderiv_within 𝕜 m.succ f s x, - rw ← iterated_fderiv_within_inter (is_open.mem_nhds o_open xo) hs hx, - exact (Hp.mono ho).eq_ftaylor_series_of_unique_diff_on le_rfl - (hs.inter o_open) ⟨hx, xo⟩ }, - rw [← this, ← has_fderiv_within_at_inter (is_open.mem_nhds o_open xo)], - have A : ∀ y ∈ s ∩ o, p y m = ftaylor_series_within 𝕜 f s y m, - { rintros y ⟨hy, yo⟩, - change p y m = iterated_fderiv_within 𝕜 m f s y, - rw ← iterated_fderiv_within_inter (is_open.mem_nhds o_open yo) hs hy, - exact (Hp.mono ho).eq_ftaylor_series_of_unique_diff_on (with_top.coe_le_coe.2 (nat.le_succ m)) - (hs.inter o_open) ⟨hy, yo⟩ }, - exact ((Hp.mono ho).fderiv_within m (with_top.coe_lt_coe.2 (lt_add_one m)) x ⟨hx, xo⟩).congr - (λ y hy, (A y hy).symm) (A x ⟨hx, xo⟩).symm }, - { assume m hm, - apply continuous_on_of_locally_continuous_on, - assume x hx, - rcases h x hx m hm with ⟨u, hu, p, Hp⟩, - rcases mem_nhds_within.1 hu with ⟨o, o_open, xo, ho⟩, - rw insert_eq_of_mem hx at ho, - rw inter_comm at ho, - refine ⟨o, o_open, xo, _⟩, - have A : ∀ y ∈ s ∩ o, p y m = ftaylor_series_within 𝕜 f s y m, - { rintros y ⟨hy, yo⟩, - change p y m = iterated_fderiv_within 𝕜 m f s y, - rw ← iterated_fderiv_within_inter (is_open.mem_nhds o_open yo) hs hy, - exact (Hp.mono ho).eq_ftaylor_series_of_unique_diff_on le_rfl - (hs.inter o_open) ⟨hy, yo⟩ }, - exact ((Hp.mono ho).cont m le_rfl).congr (λ y hy, (A y hy).symm) } -end - -lemma cont_diff_on_of_continuous_on_differentiable_on - (Hcont : ∀ (m : ℕ), (m : ℕ∞) ≤ n → - continuous_on (λ x, iterated_fderiv_within 𝕜 m f s x) s) - (Hdiff : ∀ (m : ℕ), (m : ℕ∞) < n → - differentiable_on 𝕜 (λ x, iterated_fderiv_within 𝕜 m f s x) s) : - cont_diff_on 𝕜 n f s := -begin - assume x hx m hm, - rw insert_eq_of_mem hx, - refine ⟨s, self_mem_nhds_within, ftaylor_series_within 𝕜 f s, _⟩, - split, - { assume y hy, - simp only [ftaylor_series_within, continuous_multilinear_map.uncurry0_apply, - iterated_fderiv_within_zero_apply] }, - { assume k hk y hy, - convert (Hdiff k (lt_of_lt_of_le hk hm) y hy).has_fderiv_within_at, - simp only [ftaylor_series_within, iterated_fderiv_within_succ_eq_comp_left, - continuous_linear_equiv.coe_apply, function.comp_app, coe_fn_coe_base], - exact continuous_linear_map.curry_uncurry_left _ }, - { assume k hk, - exact Hcont k (le_trans hk hm) } -end - -lemma cont_diff_on_of_differentiable_on - (h : ∀(m : ℕ), (m : ℕ∞) ≤ n → differentiable_on 𝕜 (iterated_fderiv_within 𝕜 m f s) s) : - cont_diff_on 𝕜 n f s := -cont_diff_on_of_continuous_on_differentiable_on - (λ m hm, (h m hm).continuous_on) (λ m hm, (h m (le_of_lt hm))) - -lemma cont_diff_on.continuous_on_iterated_fderiv_within {m : ℕ} - (h : cont_diff_on 𝕜 n f s) (hmn : (m : ℕ∞) ≤ n) (hs : unique_diff_on 𝕜 s) : - continuous_on (iterated_fderiv_within 𝕜 m f s) s := -(h.ftaylor_series_within hs).cont m hmn - -lemma cont_diff_on.differentiable_on_iterated_fderiv_within {m : ℕ} - (h : cont_diff_on 𝕜 n f s) (hmn : (m : ℕ∞) < n) (hs : unique_diff_on 𝕜 s) : - differentiable_on 𝕜 (iterated_fderiv_within 𝕜 m f s) s := -λ x hx, ((h.ftaylor_series_within hs).fderiv_within m hmn x hx).differentiable_within_at - -lemma cont_diff_on_iff_continuous_on_differentiable_on - (hs : unique_diff_on 𝕜 s) : - cont_diff_on 𝕜 n f s ↔ - (∀ (m : ℕ), (m : ℕ∞) ≤ n → - continuous_on (λ x, iterated_fderiv_within 𝕜 m f s x) s) - ∧ (∀ (m : ℕ), (m : ℕ∞) < n → - differentiable_on 𝕜 (λ x, iterated_fderiv_within 𝕜 m f s x) s) := -begin - split, - { assume h, - split, - { assume m hm, exact h.continuous_on_iterated_fderiv_within hm hs }, - { assume m hm, exact h.differentiable_on_iterated_fderiv_within hm hs } }, - { assume h, - 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 - 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 -differentiable there, and its derivative (expressed with `fderiv`) is `C^n`. -/ -theorem cont_diff_on_succ_iff_fderiv_of_open {n : ℕ} (hs : is_open s) : - cont_diff_on 𝕜 ((n + 1) : ℕ) f s ↔ - differentiable_on 𝕜 f s ∧ cont_diff_on 𝕜 n (λ y, fderiv 𝕜 f y) s := -begin - rw cont_diff_on_succ_iff_fderiv_within hs.unique_diff_on, - congrm _ ∧ _, - apply cont_diff_on_congr, - assume x hx, - exact fderiv_within_of_open hs hx -end - -/-- A function is `C^∞` on a domain with unique derivatives if and only if it is differentiable -there, and its derivative (expressed with `fderiv_within`) is `C^∞`. -/ -theorem cont_diff_on_top_iff_fderiv_within (hs : unique_diff_on 𝕜 s) : - cont_diff_on 𝕜 ∞ f s ↔ - differentiable_on 𝕜 f s ∧ cont_diff_on 𝕜 ∞ (λ y, fderiv_within 𝕜 f s y) s := -begin - split, - { assume h, - refine ⟨h.differentiable_on le_top, _⟩, - apply cont_diff_on_top.2 (λ n, ((cont_diff_on_succ_iff_fderiv_within hs).1 _).2), - exact h.of_le le_top }, - { assume h, - refine cont_diff_on_top.2 (λ n, _), - have A : (n : ℕ∞) ≤ ∞ := le_top, - apply ((cont_diff_on_succ_iff_fderiv_within hs).2 ⟨h.1, h.2.of_le A⟩).of_le, - exact with_top.coe_le_coe.2 (nat.le_succ n) } -end - -/-- A function is `C^∞` on an open domain if and only if it is differentiable there, and its -derivative (expressed with `fderiv`) is `C^∞`. -/ -theorem cont_diff_on_top_iff_fderiv_of_open (hs : is_open s) : - cont_diff_on 𝕜 ∞ f s ↔ - differentiable_on 𝕜 f s ∧ cont_diff_on 𝕜 ∞ (λ y, fderiv 𝕜 f y) s := -begin - rw cont_diff_on_top_iff_fderiv_within hs.unique_diff_on, - congrm _ ∧ _, - apply cont_diff_on_congr, - assume x hx, - exact fderiv_within_of_open hs hx -end - -lemma cont_diff_on.fderiv_within - (hf : cont_diff_on 𝕜 n f s) (hs : unique_diff_on 𝕜 s) (hmn : m + 1 ≤ n) : - cont_diff_on 𝕜 m (λ y, fderiv_within 𝕜 f s y) s := -begin - cases m, - { change ∞ + 1 ≤ n at hmn, - have : n = ∞, by simpa using hmn, - rw this at hf, - exact ((cont_diff_on_top_iff_fderiv_within hs).1 hf).2 }, - { change (m.succ : ℕ∞) ≤ n at hmn, - exact ((cont_diff_on_succ_iff_fderiv_within hs).1 (hf.of_le hmn)).2 } -end - -lemma cont_diff_on.fderiv_of_open - (hf : cont_diff_on 𝕜 n f s) (hs : is_open s) (hmn : m + 1 ≤ n) : - cont_diff_on 𝕜 m (λ y, fderiv 𝕜 f y) s := -(hf.fderiv_within hs.unique_diff_on hmn).congr (λ x hx, (fderiv_within_of_open hs hx).symm) - -lemma cont_diff_on.continuous_on_fderiv_within - (h : cont_diff_on 𝕜 n f s) (hs : unique_diff_on 𝕜 s) (hn : 1 ≤ n) : - continuous_on (λ x, fderiv_within 𝕜 f s x) s := -((cont_diff_on_succ_iff_fderiv_within hs).1 (h.of_le hn)).2.continuous_on - -lemma cont_diff_on.continuous_on_fderiv_of_open - (h : cont_diff_on 𝕜 n f s) (hs : is_open s) (hn : 1 ≤ n) : - continuous_on (λ x, fderiv 𝕜 f x) s := -((cont_diff_on_succ_iff_fderiv_of_open hs).1 (h.of_le hn)).2.continuous_on - -/-! ### Functions with a Taylor series on the whole space -/ - -/-- `has_ftaylor_series_up_to n f p` registers the fact that `p 0 = f` and `p (m+1)` is a -derivative of `p m` for `m < n`, and is continuous for `m ≤ n`. This is a predicate analogous to -`has_fderiv_at` but for higher order derivatives. -/ -structure has_ftaylor_series_up_to (n : ℕ∞) - (f : E → F) (p : E → formal_multilinear_series 𝕜 E F) : Prop := -(zero_eq : ∀ x, (p x 0).uncurry0 = f x) -(fderiv : ∀ (m : ℕ) (hm : (m : ℕ∞) < n), ∀ x, - has_fderiv_at (λ y, p y m) (p x m.succ).curry_left x) -(cont : ∀ (m : ℕ) (hm : (m : ℕ∞) ≤ n), continuous (λ x, p x m)) - -lemma has_ftaylor_series_up_to.zero_eq' - (h : has_ftaylor_series_up_to n f p) (x : E) : - p x 0 = (continuous_multilinear_curry_fin0 𝕜 E F).symm (f x) := -by { rw ← h.zero_eq x, symmetry, exact continuous_multilinear_map.uncurry0_curry0 _ } - -lemma has_ftaylor_series_up_to_on_univ_iff : - has_ftaylor_series_up_to_on n f p univ ↔ has_ftaylor_series_up_to n f p := -begin - split, - { assume H, - split, - { exact λ x, H.zero_eq x (mem_univ x) }, - { assume m hm x, - rw ← has_fderiv_within_at_univ, - exact H.fderiv_within m hm x (mem_univ x) }, - { assume m hm, - rw continuous_iff_continuous_on_univ, - exact H.cont m hm } }, - { assume H, - split, - { exact λ x hx, H.zero_eq x }, - { assume m hm x hx, - rw has_fderiv_within_at_univ, - exact H.fderiv m hm x }, - { assume m hm, - rw ← continuous_iff_continuous_on_univ, - exact H.cont m hm } } -end - -lemma has_ftaylor_series_up_to.has_ftaylor_series_up_to_on - (h : has_ftaylor_series_up_to n f p) (s : set E) : - has_ftaylor_series_up_to_on n f p s := -(has_ftaylor_series_up_to_on_univ_iff.2 h).mono (subset_univ _) - -lemma has_ftaylor_series_up_to.of_le - (h : has_ftaylor_series_up_to n f p) (hmn : m ≤ n) : - has_ftaylor_series_up_to m f p := -by { rw ← has_ftaylor_series_up_to_on_univ_iff at h ⊢, exact h.of_le hmn } - -lemma has_ftaylor_series_up_to.continuous - (h : has_ftaylor_series_up_to n f p) : continuous f := -begin - rw ← has_ftaylor_series_up_to_on_univ_iff at h, - rw continuous_iff_continuous_on_univ, - exact h.continuous_on -end - -lemma has_ftaylor_series_up_to_zero_iff : - has_ftaylor_series_up_to 0 f p ↔ continuous f ∧ (∀ x, (p x 0).uncurry0 = f x) := -by simp [has_ftaylor_series_up_to_on_univ_iff.symm, continuous_iff_continuous_on_univ, - has_ftaylor_series_up_to_on_zero_iff] - -/-- If a function has a Taylor series at order at least `1`, then the term of order `1` of this -series is a derivative of `f`. -/ -lemma has_ftaylor_series_up_to.has_fderiv_at - (h : has_ftaylor_series_up_to n f p) (hn : 1 ≤ n) (x : E) : - has_fderiv_at f (continuous_multilinear_curry_fin1 𝕜 E F (p x 1)) x := -begin - rw [← has_fderiv_within_at_univ], - exact (has_ftaylor_series_up_to_on_univ_iff.2 h).has_fderiv_within_at hn (mem_univ _) -end - -lemma has_ftaylor_series_up_to.differentiable - (h : has_ftaylor_series_up_to n f p) (hn : 1 ≤ n) : differentiable 𝕜 f := -λ x, (h.has_fderiv_at hn x).differentiable_at - -/-- `p` is a Taylor series of `f` up to `n+1` if and only if `p.shift` is a Taylor series up to `n` -for `p 1`, which is a derivative of `f`. -/ -theorem has_ftaylor_series_up_to_succ_iff_right {n : ℕ} : - has_ftaylor_series_up_to ((n + 1) : ℕ) f p ↔ - (∀ x, (p x 0).uncurry0 = f x) - ∧ (∀ x, has_fderiv_at (λ y, p y 0) (p x 1).curry_left x) - ∧ has_ftaylor_series_up_to n - (λ x, continuous_multilinear_curry_fin1 𝕜 E F (p x 1)) (λ x, (p x).shift) := -by simp only [has_ftaylor_series_up_to_on_succ_iff_right, ← has_ftaylor_series_up_to_on_univ_iff, - mem_univ, forall_true_left, has_fderiv_within_at_univ] - -/-! ### Smooth functions at a point -/ - -variable (𝕜) - -/-- A function is continuously differentiable up to `n` at a point `x` if, for any integer `k ≤ n`, -there is a neighborhood of `x` where `f` admits derivatives up to order `n`, which are continuous. --/ -def cont_diff_at (n : ℕ∞) (f : E → F) (x : E) : Prop := -cont_diff_within_at 𝕜 n f univ x - -variable {𝕜} - -theorem cont_diff_within_at_univ : - cont_diff_within_at 𝕜 n f univ x ↔ cont_diff_at 𝕜 n f x := -iff.rfl - -lemma cont_diff_at_top : - cont_diff_at 𝕜 ∞ f x ↔ ∀ (n : ℕ), cont_diff_at 𝕜 n f x := -by simp [← cont_diff_within_at_univ, cont_diff_within_at_top] - -lemma cont_diff_at.cont_diff_within_at - (h : cont_diff_at 𝕜 n f x) : cont_diff_within_at 𝕜 n f s x := -h.mono (subset_univ _) - -lemma cont_diff_within_at.cont_diff_at - (h : cont_diff_within_at 𝕜 n f s x) (hx : s ∈ 𝓝 x) : - cont_diff_at 𝕜 n f x := -by rwa [cont_diff_at, ← cont_diff_within_at_inter hx, univ_inter] - -lemma cont_diff_at.congr_of_eventually_eq - (h : cont_diff_at 𝕜 n f x) (hg : f₁ =ᶠ[𝓝 x] f) : - cont_diff_at 𝕜 n f₁ x := -h.congr_of_eventually_eq' (by rwa nhds_within_univ) (mem_univ x) - -lemma cont_diff_at.of_le - (h : cont_diff_at 𝕜 n f x) (hmn : m ≤ n) : - cont_diff_at 𝕜 m f x := -h.of_le hmn - -lemma cont_diff_at.continuous_at - (h : cont_diff_at 𝕜 n f x) : continuous_at f x := -by simpa [continuous_within_at_univ] using h.continuous_within_at - -/-- If a function is `C^n` with `n ≥ 1` at a point, then it is differentiable there. -/ -lemma cont_diff_at.differentiable_at - (h : cont_diff_at 𝕜 n f x) (hn : 1 ≤ n) : differentiable_at 𝕜 f x := -by simpa [hn, differentiable_within_at_univ] using h.differentiable_within_at - -/-- A function is `C^(n + 1)` at a point iff locally, it has a derivative which is `C^n`. -/ -theorem cont_diff_at_succ_iff_has_fderiv_at {n : ℕ} : - cont_diff_at 𝕜 ((n + 1) : ℕ) f x - ↔ (∃ f' : E → E →L[𝕜] F, (∃ u ∈ 𝓝 x, ∀ x ∈ u, has_fderiv_at f (f' x) x) - ∧ cont_diff_at 𝕜 n f' x) := -begin - rw [← cont_diff_within_at_univ, cont_diff_within_at_succ_iff_has_fderiv_within_at], - simp only [nhds_within_univ, exists_prop, mem_univ, insert_eq_of_mem], - split, - { rintros ⟨u, H, f', h_fderiv, h_cont_diff⟩, - rcases mem_nhds_iff.mp H with ⟨t, htu, ht, hxt⟩, - refine ⟨f', ⟨t, _⟩, h_cont_diff.cont_diff_at H⟩, - refine ⟨mem_nhds_iff.mpr ⟨t, subset.rfl, ht, hxt⟩, _⟩, - intros y hyt, - refine (h_fderiv y (htu hyt)).has_fderiv_at _, - exact mem_nhds_iff.mpr ⟨t, htu, ht, hyt⟩ }, - { rintros ⟨f', ⟨u, H, h_fderiv⟩, h_cont_diff⟩, - refine ⟨u, H, f', _, h_cont_diff.cont_diff_within_at⟩, - intros x hxu, - exact (h_fderiv x hxu).has_fderiv_within_at } -end - -protected theorem cont_diff_at.eventually {n : ℕ} (h : cont_diff_at 𝕜 n f x) : - ∀ᶠ y in 𝓝 x, cont_diff_at 𝕜 n f y := -by simpa [nhds_within_univ] using h.eventually - -/-! ### Smooth functions -/ - -variable (𝕜) - -/-- A function is continuously differentiable up to `n` if it admits derivatives up to -order `n`, which are continuous. Contrary to the case of definitions in domains (where derivatives -might not be unique) we do not need to localize the definition in space or time. --/ -def cont_diff (n : ℕ∞) (f : E → F) : Prop := -∃ p : E → formal_multilinear_series 𝕜 E F, has_ftaylor_series_up_to n f p - -variable {𝕜} - -theorem cont_diff_on_univ : cont_diff_on 𝕜 n f univ ↔ cont_diff 𝕜 n f := -begin - split, - { assume H, - use ftaylor_series_within 𝕜 f univ, - rw ← has_ftaylor_series_up_to_on_univ_iff, - exact H.ftaylor_series_within unique_diff_on_univ }, - { rintros ⟨p, hp⟩ x hx m hm, - exact ⟨univ, filter.univ_sets _, p, (hp.has_ftaylor_series_up_to_on univ).of_le hm⟩ } -end - -lemma cont_diff_iff_cont_diff_at : cont_diff 𝕜 n f ↔ ∀ x, cont_diff_at 𝕜 n f x := -by simp [← cont_diff_on_univ, cont_diff_on, cont_diff_at] - -lemma cont_diff.cont_diff_at (h : cont_diff 𝕜 n f) : cont_diff_at 𝕜 n f x := -cont_diff_iff_cont_diff_at.1 h x - -lemma cont_diff.cont_diff_within_at (h : cont_diff 𝕜 n f) : cont_diff_within_at 𝕜 n f s x := -h.cont_diff_at.cont_diff_within_at - -lemma cont_diff_top : cont_diff 𝕜 ∞ f ↔ ∀ (n : ℕ), cont_diff 𝕜 n f := -by simp [cont_diff_on_univ.symm, cont_diff_on_top] - -lemma cont_diff_all_iff_nat : (∀ n, cont_diff 𝕜 n f) ↔ (∀ n : ℕ, cont_diff 𝕜 n f) := -by simp only [← cont_diff_on_univ, cont_diff_on_all_iff_nat] - -lemma cont_diff.cont_diff_on (h : cont_diff 𝕜 n f) : cont_diff_on 𝕜 n f s := -(cont_diff_on_univ.2 h).mono (subset_univ _) - -@[simp] lemma cont_diff_zero : cont_diff 𝕜 0 f ↔ continuous f := -begin - rw [← cont_diff_on_univ, continuous_iff_continuous_on_univ], - exact cont_diff_on_zero -end - -lemma cont_diff_at_zero : cont_diff_at 𝕜 0 f x ↔ ∃ u ∈ 𝓝 x, continuous_on f u := -by { rw ← cont_diff_within_at_univ, simp [cont_diff_within_at_zero, nhds_within_univ] } - -theorem cont_diff_at_one_iff : cont_diff_at 𝕜 1 f x ↔ - ∃ f' : E → (E →L[𝕜] F), ∃ u ∈ 𝓝 x, continuous_on f' u ∧ ∀ x ∈ u, has_fderiv_at f (f' x) x := -by simp_rw [show (1 : ℕ∞) = (0 + 1 : ℕ), from (zero_add 1).symm, - cont_diff_at_succ_iff_has_fderiv_at, show ((0 : ℕ) : ℕ∞) = 0, from rfl, - cont_diff_at_zero, exists_mem_and_iff antitone_bforall antitone_continuous_on, and_comm] - -lemma cont_diff.of_le (h : cont_diff 𝕜 n f) (hmn : m ≤ n) : cont_diff 𝕜 m f := -cont_diff_on_univ.1 $ (cont_diff_on_univ.2 h).of_le hmn - -lemma cont_diff.of_succ {n : ℕ} (h : cont_diff 𝕜 (n + 1) f) : cont_diff 𝕜 n f := -h.of_le $ with_top.coe_le_coe.mpr le_self_add - -lemma cont_diff.one_of_succ {n : ℕ} (h : cont_diff 𝕜 (n + 1) f) : cont_diff 𝕜 1 f := -h.of_le $ with_top.coe_le_coe.mpr le_add_self - -lemma cont_diff.continuous (h : cont_diff 𝕜 n f) : continuous f := -cont_diff_zero.1 (h.of_le bot_le) - -/-- If a function is `C^n` with `n ≥ 1`, then it is differentiable. -/ -lemma cont_diff.differentiable (h : cont_diff 𝕜 n f) (hn : 1 ≤ n) : differentiable 𝕜 f := -differentiable_on_univ.1 $ (cont_diff_on_univ.2 h).differentiable_on hn - -lemma cont_diff_iff_forall_nat_le : - cont_diff 𝕜 n f ↔ ∀ m : ℕ, ↑m ≤ n → cont_diff 𝕜 m f := -by { simp_rw [← cont_diff_on_univ], exact cont_diff_on_iff_forall_nat_le } - - -/-! ### Iterated derivative -/ - -variable (𝕜) - -/-- The `n`-th derivative of a function, as a multilinear map, defined inductively. -/ -noncomputable def iterated_fderiv (n : ℕ) (f : E → F) : - E → (E [×n]→L[𝕜] F) := -nat.rec_on n - (λ x, continuous_multilinear_map.curry0 𝕜 E (f x)) - (λ n rec x, continuous_linear_map.uncurry_left (fderiv 𝕜 rec x)) - -/-- Formal Taylor series associated to a function within a set. -/ -def ftaylor_series (f : E → F) (x : E) : formal_multilinear_series 𝕜 E F := -λ n, iterated_fderiv 𝕜 n f x - -variable {𝕜} - -@[simp] lemma iterated_fderiv_zero_apply (m : (fin 0) → E) : - (iterated_fderiv 𝕜 0 f x : ((fin 0) → E) → F) m = f x := rfl - -lemma iterated_fderiv_zero_eq_comp : - iterated_fderiv 𝕜 0 f = (continuous_multilinear_curry_fin0 𝕜 E F).symm ∘ f := rfl - -@[simp] lemma norm_iterated_fderiv_zero : - ‖iterated_fderiv 𝕜 0 f x‖ = ‖f x‖ := -by rw [iterated_fderiv_zero_eq_comp, linear_isometry_equiv.norm_map] - -lemma iterated_fderiv_with_zero_eq : - iterated_fderiv_within 𝕜 0 f s = iterated_fderiv 𝕜 0 f := -by { ext, refl } - -lemma iterated_fderiv_succ_apply_left {n : ℕ} (m : fin (n + 1) → E): - (iterated_fderiv 𝕜 (n + 1) f x : (fin (n + 1) → E) → F) m - = (fderiv 𝕜 (iterated_fderiv 𝕜 n f) x : E → (E [×n]→L[𝕜] F)) (m 0) (tail m) := rfl - -/-- Writing explicitly the `n+1`-th derivative as the composition of a currying linear equiv, -and the derivative of the `n`-th derivative. -/ -lemma iterated_fderiv_succ_eq_comp_left {n : ℕ} : - iterated_fderiv 𝕜 (n + 1) f = - (continuous_multilinear_curry_left_equiv 𝕜 (λ (i : fin (n + 1)), E) F) - ∘ (fderiv 𝕜 (iterated_fderiv 𝕜 n f)) := rfl - -/-- Writing explicitly the derivative of the `n`-th derivative as the composition of a currying -linear equiv, and the `n + 1`-th derivative. -/ -lemma fderiv_iterated_fderiv {n : ℕ} : - fderiv 𝕜 (iterated_fderiv 𝕜 n f) = - (continuous_multilinear_curry_left_equiv 𝕜 (λ (i : fin (n + 1)), E) F).symm - ∘ (iterated_fderiv 𝕜 (n + 1) f) := -begin - rw iterated_fderiv_succ_eq_comp_left, - ext1 x, - simp only [function.comp_app, linear_isometry_equiv.symm_apply_apply], -end - -lemma has_compact_support.iterated_fderiv (hf : has_compact_support f) (n : ℕ) : - has_compact_support (iterated_fderiv 𝕜 n f) := -begin - induction n with n IH, - { rw [iterated_fderiv_zero_eq_comp], - apply hf.comp_left, - exact linear_isometry_equiv.map_zero _ }, - { rw iterated_fderiv_succ_eq_comp_left, - apply (IH.fderiv 𝕜).comp_left, - exact linear_isometry_equiv.map_zero _ } -end -lemma norm_fderiv_iterated_fderiv {n : ℕ} : - ‖fderiv 𝕜 (iterated_fderiv 𝕜 n f) x‖ = ‖iterated_fderiv 𝕜 (n + 1) f x‖ := -by rw [iterated_fderiv_succ_eq_comp_left, linear_isometry_equiv.norm_map] - -lemma iterated_fderiv_within_univ {n : ℕ} : - iterated_fderiv_within 𝕜 n f univ = iterated_fderiv 𝕜 n f := -begin - induction n with n IH, - { ext x, simp }, - { ext x m, - rw [iterated_fderiv_succ_apply_left, iterated_fderiv_within_succ_apply_left, IH, - fderiv_within_univ] } -end - -/-- In an open set, the iterated derivative within this set coincides with the global iterated -derivative. -/ -lemma iterated_fderiv_within_of_is_open (n : ℕ) (hs : is_open s) : - eq_on (iterated_fderiv_within 𝕜 n f s) (iterated_fderiv 𝕜 n f) s := -begin - induction n with n IH, - { assume x hx, - ext1 m, - simp only [iterated_fderiv_within_zero_apply, iterated_fderiv_zero_apply] }, - { assume x hx, - rw [iterated_fderiv_succ_eq_comp_left, iterated_fderiv_within_succ_eq_comp_left], - dsimp, - congr' 1, - rw fderiv_within_of_open hs hx, - apply filter.eventually_eq.fderiv_eq, - filter_upwards [hs.mem_nhds hx], - exact IH } -end - -lemma ftaylor_series_within_univ : - ftaylor_series_within 𝕜 f univ = ftaylor_series 𝕜 f := -begin - ext1 x, ext1 n, - change iterated_fderiv_within 𝕜 n f univ x = iterated_fderiv 𝕜 n f x, - rw iterated_fderiv_within_univ -end - -theorem iterated_fderiv_succ_apply_right {n : ℕ} (m : fin (n + 1) → E) : - (iterated_fderiv 𝕜 (n + 1) f x : (fin (n + 1) → E) → F) m - = iterated_fderiv 𝕜 n (λy, fderiv 𝕜 f y) x (init m) (m (last n)) := -begin - rw [← iterated_fderiv_within_univ, ← iterated_fderiv_within_univ, ← fderiv_within_univ], - exact iterated_fderiv_within_succ_apply_right unique_diff_on_univ (mem_univ _) _ -end - -/-- Writing explicitly the `n+1`-th derivative as the composition of a currying linear equiv, -and the `n`-th derivative of the derivative. -/ -lemma iterated_fderiv_succ_eq_comp_right {n : ℕ} : - iterated_fderiv 𝕜 (n + 1) f x = - ((continuous_multilinear_curry_right_equiv' 𝕜 n E F) - ∘ (iterated_fderiv 𝕜 n (λy, fderiv 𝕜 f y))) x := -by { ext m, rw iterated_fderiv_succ_apply_right, refl } - -lemma norm_iterated_fderiv_fderiv {n : ℕ} : - ‖iterated_fderiv 𝕜 n (fderiv 𝕜 f) x‖ = ‖iterated_fderiv 𝕜 (n + 1) f x‖ := -by rw [iterated_fderiv_succ_eq_comp_right, linear_isometry_equiv.norm_map] - -@[simp] lemma iterated_fderiv_one_apply (m : (fin 1) → E) : - (iterated_fderiv 𝕜 1 f x : ((fin 1) → E) → F) m - = (fderiv 𝕜 f x : E → F) (m 0) := -by { rw [iterated_fderiv_succ_apply_right, iterated_fderiv_zero_apply], refl } - -/-- When a function is `C^n` in a set `s` of unique differentiability, it admits -`ftaylor_series_within 𝕜 f s` as a Taylor series up to order `n` in `s`. -/ -theorem cont_diff_iff_ftaylor_series : - cont_diff 𝕜 n f ↔ has_ftaylor_series_up_to n f (ftaylor_series 𝕜 f) := -begin - split, - { rw [← cont_diff_on_univ, ← has_ftaylor_series_up_to_on_univ_iff, - ← ftaylor_series_within_univ], - exact λ h, cont_diff_on.ftaylor_series_within h unique_diff_on_univ }, - { assume h, exact ⟨ftaylor_series 𝕜 f, h⟩ } -end - -lemma cont_diff_iff_continuous_differentiable : - cont_diff 𝕜 n f ↔ - (∀ (m : ℕ), (m : ℕ∞) ≤ n → continuous (λ x, iterated_fderiv 𝕜 m f x)) - ∧ (∀ (m : ℕ), (m : ℕ∞) < n → differentiable 𝕜 (λ x, iterated_fderiv 𝕜 m f x)) := -by simp [cont_diff_on_univ.symm, continuous_iff_continuous_on_univ, - differentiable_on_univ.symm, iterated_fderiv_within_univ, - cont_diff_on_iff_continuous_on_differentiable_on unique_diff_on_univ] - -/-- If `f` is `C^n` then its `m`-times iterated derivative is continuous for `m ≤ n`. -/ -lemma cont_diff.continuous_iterated_fderiv {m : ℕ} (hm : (m : ℕ∞) ≤ n) - (hf : cont_diff 𝕜 n f) : continuous (λ x, iterated_fderiv 𝕜 m f x) := -(cont_diff_iff_continuous_differentiable.mp hf).1 m hm - -/-- If `f` is `C^n` then its `m`-times iterated derivative is differentiable for `m < n`. -/ -lemma cont_diff.differentiable_iterated_fderiv {m : ℕ} (hm : (m : ℕ∞) < n) - (hf : cont_diff 𝕜 n f) : differentiable 𝕜 (λ x, iterated_fderiv 𝕜 m f x) := -(cont_diff_iff_continuous_differentiable.mp hf).2 m hm - -lemma cont_diff_of_differentiable_iterated_fderiv - (h : ∀(m : ℕ), (m : ℕ∞) ≤ n → differentiable 𝕜 (iterated_fderiv 𝕜 m f)) : - cont_diff 𝕜 n f := -cont_diff_iff_continuous_differentiable.2 -⟨λ m hm, (h m hm).continuous, λ m hm, (h m (le_of_lt hm))⟩ - -/-- A function is `C^(n + 1)` if and only if it is differentiable, -and its derivative (formulated in terms of `fderiv`) is `C^n`. -/ -theorem cont_diff_succ_iff_fderiv {n : ℕ} : - cont_diff 𝕜 ((n + 1) : ℕ) f ↔ - differentiable 𝕜 f ∧ cont_diff 𝕜 n (λ y, fderiv 𝕜 f y) := -by simp only [← cont_diff_on_univ, ← differentiable_on_univ, ← fderiv_within_univ, - cont_diff_on_succ_iff_fderiv_within unique_diff_on_univ] - -theorem cont_diff_one_iff_fderiv : - cont_diff 𝕜 1 f ↔ differentiable 𝕜 f ∧ continuous (fderiv 𝕜 f) := -cont_diff_succ_iff_fderiv.trans $ iff.rfl.and cont_diff_zero - -/-- A function is `C^∞` if and only if it is differentiable, -and its derivative (formulated in terms of `fderiv`) is `C^∞`. -/ -theorem cont_diff_top_iff_fderiv : - cont_diff 𝕜 ∞ f ↔ - differentiable 𝕜 f ∧ cont_diff 𝕜 ∞ (λ y, fderiv 𝕜 f y) := -begin - simp only [← cont_diff_on_univ, ← differentiable_on_univ, ← fderiv_within_univ], - rw cont_diff_on_top_iff_fderiv_within unique_diff_on_univ, -end - -lemma cont_diff.continuous_fderiv - (h : cont_diff 𝕜 n f) (hn : 1 ≤ n) : - continuous (λ x, fderiv 𝕜 f x) := -((cont_diff_succ_iff_fderiv).1 (h.of_le hn)).2.continuous - -/-- If a function is at least `C^1`, its bundled derivative (mapping `(x, v)` to `Df(x) v`) is -continuous. -/ -lemma cont_diff.continuous_fderiv_apply - (h : cont_diff 𝕜 n f) (hn : 1 ≤ n) : - continuous (λp : E × E, (fderiv 𝕜 f p.1 : E → F) p.2) := -have A : continuous (λq : (E →L[𝕜] F) × E, q.1 q.2) := is_bounded_bilinear_map_apply.continuous, -have B : continuous (λp : E × E, (fderiv 𝕜 f p.1, p.2)) := - ((h.continuous_fderiv hn).comp continuous_fst).prod_mk continuous_snd, -A.comp B +{b : E × F → G} {m n : ℕ∞} {p : E → formal_multilinear_series 𝕜 E F} /-! ### Constants -/ diff --git a/src/analysis/calculus/cont_diff_def.lean b/src/analysis/calculus/cont_diff_def.lean new file mode 100644 index 0000000000000..dc54c755ea169 --- /dev/null +++ b/src/analysis/calculus/cont_diff_def.lean @@ -0,0 +1,1595 @@ +/- +Copyright (c) 2019 Sébastien Gouëzel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sébastien Gouëzel +-/ +import analysis.calculus.fderiv +import analysis.normed_space.multilinear +import analysis.calculus.formal_multilinear_series +import data.enat.basic + +/-! +# Higher differentiability + +A function is `C^1` on a domain if it is differentiable there, and its derivative is continuous. +By induction, it is `C^n` if it is `C^{n-1}` and its (n-1)-th derivative is `C^1` there or, +equivalently, if it is `C^1` and its derivative is `C^{n-1}`. +Finally, it is `C^∞` if it is `C^n` for all n. + +We formalize these notions by defining iteratively the `n+1`-th derivative of a function as the +derivative of the `n`-th derivative. It is called `iterated_fderiv 𝕜 n f x` where `𝕜` is the +field, `n` is the number of iterations, `f` is the function and `x` is the point, and it is given +as an `n`-multilinear map. We also define a version `iterated_fderiv_within` relative to a domain, +as well as predicates `cont_diff_within_at`, `cont_diff_at`, `cont_diff_on` and +`cont_diff` saying that the function is `C^n` within a set at a point, at a point, on a set +and on the whole space respectively. + +To avoid the issue of choice when choosing a derivative in sets where the derivative is not +necessarily unique, `cont_diff_on` is not defined directly in terms of the +regularity of the specific choice `iterated_fderiv_within 𝕜 n f s` inside `s`, but in terms of the +existence of a nice sequence of derivatives, expressed with a predicate +`has_ftaylor_series_up_to_on`. + +We prove basic properties of these notions. + +## Main definitions and results +Let `f : E → F` be a map between normed vector spaces over a nontrivially normed field `𝕜`. + +* `has_ftaylor_series_up_to n f p`: expresses that the formal multilinear series `p` is a sequence + of iterated derivatives of `f`, up to the `n`-th term (where `n` is a natural number or `∞`). +* `has_ftaylor_series_up_to_on n f p s`: same thing, but inside a set `s`. The notion of derivative + is now taken inside `s`. In particular, derivatives don't have to be unique. +* `cont_diff 𝕜 n f`: expresses that `f` is `C^n`, i.e., it admits a Taylor series up to + rank `n`. +* `cont_diff_on 𝕜 n f s`: expresses that `f` is `C^n` in `s`. +* `cont_diff_at 𝕜 n f x`: expresses that `f` is `C^n` around `x`. +* `cont_diff_within_at 𝕜 n f s x`: expresses that `f` is `C^n` around `x` within the set `s`. +* `iterated_fderiv_within 𝕜 n f s x` is an `n`-th derivative of `f` over the field `𝕜` on the + set `s` at the point `x`. It is a continuous multilinear map from `E^n` to `F`, defined as a + derivative within `s` of `iterated_fderiv_within 𝕜 (n-1) f s` if one exists, and `0` otherwise. +* `iterated_fderiv 𝕜 n f x` is the `n`-th derivative of `f` over the field `𝕜` at the point `x`. + It is a continuous multilinear map from `E^n` to `F`, defined as a derivative of + `iterated_fderiv 𝕜 (n-1) f` if one exists, and `0` otherwise. + +In sets of unique differentiability, `cont_diff_on 𝕜 n f s` can be expressed in terms of the +properties of `iterated_fderiv_within 𝕜 m f s` for `m ≤ n`. In the whole space, +`cont_diff 𝕜 n f` can be expressed in terms of the properties of `iterated_fderiv 𝕜 m f` +for `m ≤ n`. + +## Implementation notes + +The definitions in this file are designed to work on any field `𝕜`. They are sometimes slightly more +complicated than the naive definitions one would guess from the intuition over the real or complex +numbers, but they are designed to circumvent the lack of gluing properties and partitions of unity +in general. In the usual situations, they coincide with the usual definitions. + +### Definition of `C^n` functions in domains + +One could define `C^n` functions in a domain `s` by fixing an arbitrary choice of derivatives (this +is what we do with `iterated_fderiv_within`) and requiring that all these derivatives up to `n` are +continuous. If the derivative is not unique, this could lead to strange behavior like two `C^n` +functions `f` and `g` on `s` whose sum is not `C^n`. A better definition is thus to say that a +function is `C^n` inside `s` if it admits a sequence of derivatives up to `n` inside `s`. + +This definition still has the problem that a function which is locally `C^n` would not need to +be `C^n`, as different choices of sequences of derivatives around different points might possibly +not be glued together to give a globally defined sequence of derivatives. (Note that this issue +can not happen over reals, thanks to partition of unity, but the behavior over a general field is +not so clear, and we want a definition for general fields). Also, there are locality +problems for the order parameter: one could image a function which, for each `n`, has a nice +sequence of derivatives up to order `n`, but they do not coincide for varying `n` and can therefore +not be glued to give rise to an infinite sequence of derivatives. This would give a function +which is `C^n` for all `n`, but not `C^∞`. We solve this issue by putting locality conditions +in space and order in our definition of `cont_diff_within_at` and `cont_diff_on`. +The resulting definition is slightly more complicated to work with (in fact not so much), but it +gives rise to completely satisfactory theorems. + +For instance, with this definition, a real function which is `C^m` (but not better) on `(-1/m, 1/m)` +for each natural `m` is by definition `C^∞` at `0`. + +There is another issue with the definition of `cont_diff_within_at 𝕜 n f s x`. We can +require the existence and good behavior of derivatives up to order `n` on a neighborhood of `x` +within `s`. However, this does not imply continuity or differentiability within `s` of the function +at `x` when `x` does not belong to `s`. Therefore, we require such existence and good behavior on +a neighborhood of `x` within `s ∪ {x}` (which appears as `insert x s` in this file). + +### Side of the composition, and universe issues + +With a naïve direct definition, the `n`-th derivative of a function belongs to the space +`E →L[𝕜] (E →L[𝕜] (E ... F)...)))` where there are n iterations of `E →L[𝕜]`. This space +may also be seen as the space of continuous multilinear functions on `n` copies of `E` with +values in `F`, by uncurrying. This is the point of view that is usually adopted in textbooks, +and that we also use. This means that the definition and the first proofs are slightly involved, +as one has to keep track of the uncurrying operation. The uncurrying can be done from the +left or from the right, amounting to defining the `n+1`-th derivative either as the derivative of +the `n`-th derivative, or as the `n`-th derivative of the derivative. +For proofs, it would be more convenient to use the latter approach (from the right), +as it means to prove things at the `n+1`-th step we only need to understand well enough the +derivative in `E →L[𝕜] F` (contrary to the approach from the left, where one would need to know +enough on the `n`-th derivative to deduce things on the `n+1`-th derivative). + +However, the definition from the right leads to a universe polymorphism problem: if we define +`iterated_fderiv 𝕜 (n + 1) f x = iterated_fderiv 𝕜 n (fderiv 𝕜 f) x` by induction, we need to +generalize over all spaces (as `f` and `fderiv 𝕜 f` don't take values in the same space). It is +only possible to generalize over all spaces in some fixed universe in an inductive definition. +For `f : E → F`, then `fderiv 𝕜 f` is a map `E → (E →L[𝕜] F)`. Therefore, the definition will only +work if `F` and `E →L[𝕜] F` are in the same universe. + +This issue does not appear with the definition from the left, where one does not need to generalize +over all spaces. Therefore, we use the definition from the left. This means some proofs later on +become a little bit more complicated: to prove that a function is `C^n`, the most efficient approach +is to exhibit a formula for its `n`-th derivative and prove it is continuous (contrary to the +inductive approach where one would prove smoothness statements without giving a formula for the +derivative). In the end, this approach is still satisfactory as it is good to have formulas for the +iterated derivatives in various constructions. + +One point where we depart from this explicit approach is in the proof of smoothness of a +composition: there is a formula for the `n`-th derivative of a composition (Faà di Bruno's formula), +but it is very complicated and barely usable, while the inductive proof is very simple. Thus, we +give the inductive proof. As explained above, it works by generalizing over the target space, hence +it only works well if all spaces belong to the same universe. To get the general version, we lift +things to a common universe using a trick. + +### Variables management + +The textbook definitions and proofs use various identifications and abuse of notations, for instance +when saying that the natural space in which the derivative lives, i.e., +`E →L[𝕜] (E →L[𝕜] ( ... →L[𝕜] F))`, is the same as a space of multilinear maps. When doing things +formally, we need to provide explicit maps for these identifications, and chase some diagrams to see +everything is compatible with the identifications. In particular, one needs to check that taking the +derivative and then doing the identification, or first doing the identification and then taking the +derivative, gives the same result. The key point for this is that taking the derivative commutes +with continuous linear equivalences. Therefore, we need to implement all our identifications with +continuous linear equivs. + +## Notations + +We use the notation `E [×n]→L[𝕜] F` for the space of continuous multilinear maps on `E^n` with +values in `F`. This is the space in which the `n`-th derivative of a function from `E` to `F` lives. + +In this file, we denote `⊤ : ℕ∞` with `∞`. + +## Tags + +derivative, differentiability, higher derivative, `C^n`, multilinear, Taylor series, formal series +-/ + +noncomputable theory +open_locale classical big_operators nnreal topology + +local notation `∞` := (⊤ : ℕ∞) + +local attribute [instance, priority 1001] +normed_add_comm_group.to_add_comm_group normed_space.to_module' add_comm_group.to_add_comm_monoid + +open set fin filter function + +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] +{E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] +{F : Type*} [normed_add_comm_group F] [normed_space 𝕜 F] +{G : Type*} [normed_add_comm_group G] [normed_space 𝕜 G] +{X : Type*} [normed_add_comm_group X] [normed_space 𝕜 X] +{s s₁ t u : set E} {f f₁ : E → F} {g : F → G} {x x₀ : E} {c : F} +{m n : ℕ∞} {p : E → formal_multilinear_series 𝕜 E F} + +/-! ### Functions with a Taylor series on a domain -/ + +/-- `has_ftaylor_series_up_to_on n f p s` registers the fact that `p 0 = f` and `p (m+1)` is a +derivative of `p m` for `m < n`, and is continuous for `m ≤ n`. This is a predicate analogous to +`has_fderiv_within_at` but for higher order derivatives. -/ +structure has_ftaylor_series_up_to_on (n : ℕ∞) + (f : E → F) (p : E → formal_multilinear_series 𝕜 E F) (s : set E) : Prop := +(zero_eq : ∀ x ∈ s, (p x 0).uncurry0 = f x) +(fderiv_within : ∀ (m : ℕ) (hm : (m : ℕ∞) < n), ∀ x ∈ s, + has_fderiv_within_at (λ y, p y m) (p x m.succ).curry_left s x) +(cont : ∀ (m : ℕ) (hm : (m : ℕ∞) ≤ n), continuous_on (λ x, p x m) s) + +lemma has_ftaylor_series_up_to_on.zero_eq' + (h : has_ftaylor_series_up_to_on n f p s) {x : E} (hx : x ∈ s) : + p x 0 = (continuous_multilinear_curry_fin0 𝕜 E F).symm (f x) := +by { rw ← h.zero_eq x hx, symmetry, exact continuous_multilinear_map.uncurry0_curry0 _ } + +/-- If two functions coincide on a set `s`, then a Taylor series for the first one is as well a +Taylor series for the second one. -/ +lemma has_ftaylor_series_up_to_on.congr + (h : has_ftaylor_series_up_to_on n f p s) (h₁ : ∀ x ∈ s, f₁ x = f x) : + has_ftaylor_series_up_to_on n f₁ p s := +begin + refine ⟨λ x hx, _, h.fderiv_within, h.cont⟩, + rw h₁ x hx, + exact h.zero_eq x hx +end + +lemma has_ftaylor_series_up_to_on.mono + (h : has_ftaylor_series_up_to_on n f p s) {t : set E} (hst : t ⊆ s) : + has_ftaylor_series_up_to_on n f p t := +⟨λ x hx, h.zero_eq x (hst hx), +λ m hm x hx, (h.fderiv_within m hm x (hst hx)).mono hst, +λ m hm, (h.cont m hm).mono hst⟩ + +lemma has_ftaylor_series_up_to_on.of_le + (h : has_ftaylor_series_up_to_on n f p s) (hmn : m ≤ n) : + has_ftaylor_series_up_to_on m f p s := +⟨h.zero_eq, +λ k hk x hx, h.fderiv_within k (lt_of_lt_of_le hk hmn) x hx, +λ k hk, h.cont k (le_trans hk hmn)⟩ + +lemma has_ftaylor_series_up_to_on.continuous_on + (h : has_ftaylor_series_up_to_on n f p s) : continuous_on f s := +begin + have := (h.cont 0 bot_le).congr (λ x hx, (h.zero_eq' hx).symm), + rwa linear_isometry_equiv.comp_continuous_on_iff at this +end + +lemma has_ftaylor_series_up_to_on_zero_iff : + has_ftaylor_series_up_to_on 0 f p s ↔ continuous_on f s ∧ (∀ x ∈ s, (p x 0).uncurry0 = f x) := +begin + refine ⟨λ H, ⟨H.continuous_on, H.zero_eq⟩, + λ H, ⟨H.2, λ m hm, false.elim (not_le.2 hm bot_le), _⟩⟩, + assume m hm, + obtain rfl : m = 0, by exact_mod_cast (hm.antisymm (zero_le _)), + have : ∀ x ∈ s, p x 0 = (continuous_multilinear_curry_fin0 𝕜 E F).symm (f x), + by { assume x hx, rw ← H.2 x hx, symmetry, exact continuous_multilinear_map.uncurry0_curry0 _ }, + rw [continuous_on_congr this, linear_isometry_equiv.comp_continuous_on_iff], + exact H.1 +end + +lemma has_ftaylor_series_up_to_on_top_iff : + (has_ftaylor_series_up_to_on ∞ f p s) ↔ (∀ (n : ℕ), has_ftaylor_series_up_to_on n f p s) := +begin + split, + { assume H n, exact H.of_le le_top }, + { assume H, + split, + { exact (H 0).zero_eq }, + { assume m hm, + apply (H m.succ).fderiv_within m (with_top.coe_lt_coe.2 (lt_add_one m)) }, + { assume m hm, + apply (H m).cont m le_rfl } } +end + +/-- If a function has a Taylor series at order at least `1`, then the term of order `1` of this +series is a derivative of `f`. -/ +lemma has_ftaylor_series_up_to_on.has_fderiv_within_at + (h : has_ftaylor_series_up_to_on n f p s) (hn : 1 ≤ n) (hx : x ∈ s) : + has_fderiv_within_at f (continuous_multilinear_curry_fin1 𝕜 E F (p x 1)) s x := +begin + have A : ∀ y ∈ s, f y = (continuous_multilinear_curry_fin0 𝕜 E F) (p y 0), + { assume y hy, rw ← h.zero_eq y hy, refl }, + suffices H : has_fderiv_within_at + (λ y, continuous_multilinear_curry_fin0 𝕜 E F (p y 0)) + (continuous_multilinear_curry_fin1 𝕜 E F (p x 1)) s x, + by exact H.congr A (A x hx), + rw linear_isometry_equiv.comp_has_fderiv_within_at_iff', + have : ((0 : ℕ) : ℕ∞) < n := + lt_of_lt_of_le (with_top.coe_lt_coe.2 nat.zero_lt_one) hn, + convert h.fderiv_within _ this x hx, + ext y v, + change (p x 1) (snoc 0 y) = (p x 1) (cons y v), + unfold_coes, + congr' with i, + rw unique.eq_default i, + refl +end + +lemma has_ftaylor_series_up_to_on.differentiable_on + (h : has_ftaylor_series_up_to_on n f p s) (hn : 1 ≤ n) : differentiable_on 𝕜 f s := +λ x hx, (h.has_fderiv_within_at hn hx).differentiable_within_at + +/-- If a function has a Taylor series at order at least `1` on a neighborhood of `x`, then the term +of order `1` of this series is a derivative of `f` at `x`. -/ +lemma has_ftaylor_series_up_to_on.has_fderiv_at + (h : has_ftaylor_series_up_to_on n f p s) (hn : 1 ≤ n) (hx : s ∈ 𝓝 x) : + has_fderiv_at f (continuous_multilinear_curry_fin1 𝕜 E F (p x 1)) x := +(h.has_fderiv_within_at hn (mem_of_mem_nhds hx)).has_fderiv_at hx + +/-- If a function has a Taylor series at order at least `1` on a neighborhood of `x`, then +in a neighborhood of `x`, the term of order `1` of this series is a derivative of `f`. -/ +lemma has_ftaylor_series_up_to_on.eventually_has_fderiv_at + (h : has_ftaylor_series_up_to_on n f p s) (hn : 1 ≤ n) (hx : s ∈ 𝓝 x) : + ∀ᶠ y in 𝓝 x, has_fderiv_at f (continuous_multilinear_curry_fin1 𝕜 E F (p y 1)) y := +(eventually_eventually_nhds.2 hx).mono $ λ y hy, h.has_fderiv_at hn hy + +/-- If a function has a Taylor series at order at least `1` on a neighborhood of `x`, then +it is differentiable at `x`. -/ +lemma has_ftaylor_series_up_to_on.differentiable_at + (h : has_ftaylor_series_up_to_on n f p s) (hn : 1 ≤ n) (hx : s ∈ 𝓝 x) : + differentiable_at 𝕜 f x := +(h.has_fderiv_at hn hx).differentiable_at + +/-- `p` is a Taylor series of `f` up to `n+1` if and only if `p` is a Taylor series up to `n`, and +`p (n + 1)` is a derivative of `p n`. -/ +theorem has_ftaylor_series_up_to_on_succ_iff_left {n : ℕ} : + has_ftaylor_series_up_to_on (n + 1) f p s ↔ + has_ftaylor_series_up_to_on n f p s + ∧ (∀ x ∈ s, has_fderiv_within_at (λ y, p y n) (p x n.succ).curry_left s x) + ∧ continuous_on (λ x, p x (n + 1)) s := +begin + split, + { assume h, + exact ⟨h.of_le (with_top.coe_le_coe.2 (nat.le_succ n)), + h.fderiv_within _ (with_top.coe_lt_coe.2 (lt_add_one n)), + h.cont (n + 1) le_rfl⟩ }, + { assume h, + split, + { exact h.1.zero_eq }, + { assume m hm, + by_cases h' : m < n, + { exact h.1.fderiv_within m (with_top.coe_lt_coe.2 h') }, + { have : m = n := nat.eq_of_lt_succ_of_not_lt (with_top.coe_lt_coe.1 hm) h', + rw this, + exact h.2.1 } }, + { assume m hm, + by_cases h' : m ≤ n, + { apply h.1.cont m (with_top.coe_le_coe.2 h') }, + { have : m = (n + 1) := le_antisymm (with_top.coe_le_coe.1 hm) (not_le.1 h'), + rw this, + exact h.2.2 } } } +end + +/-- `p` is a Taylor series of `f` up to `n+1` if and only if `p.shift` is a Taylor series up to `n` +for `p 1`, which is a derivative of `f`. -/ +theorem has_ftaylor_series_up_to_on_succ_iff_right {n : ℕ} : + has_ftaylor_series_up_to_on ((n + 1) : ℕ) f p s ↔ + (∀ x ∈ s, (p x 0).uncurry0 = f x) + ∧ (∀ x ∈ s, has_fderiv_within_at (λ y, p y 0) (p x 1).curry_left s x) + ∧ has_ftaylor_series_up_to_on n + (λ x, continuous_multilinear_curry_fin1 𝕜 E F (p x 1)) (λ x, (p x).shift) s := +begin + split, + { assume H, + refine ⟨H.zero_eq, H.fderiv_within 0 (with_top.coe_lt_coe.2 (nat.succ_pos n)), _⟩, + split, + { assume x hx, refl }, + { assume m (hm : (m : ℕ∞) < n) x (hx : x ∈ s), + have A : (m.succ : ℕ∞) < n.succ, + by { rw with_top.coe_lt_coe at ⊢ hm, exact nat.lt_succ_iff.mpr hm }, + change has_fderiv_within_at + ((continuous_multilinear_curry_right_equiv' 𝕜 m E F).symm + ∘ (λ (y : E), p y m.succ)) + (p x m.succ.succ).curry_right.curry_left s x, + rw linear_isometry_equiv.comp_has_fderiv_within_at_iff', + convert H.fderiv_within _ A x hx, + ext y v, + change (p x m.succ.succ) (snoc (cons y (init v)) (v (last _))) + = (p x (nat.succ (nat.succ m))) (cons y v), + rw [← cons_snoc_eq_snoc_cons, snoc_init_self] }, + { assume m (hm : (m : ℕ∞) ≤ n), + have A : (m.succ : ℕ∞) ≤ n.succ, + by { rw with_top.coe_le_coe at ⊢ hm, exact nat.pred_le_iff.mp hm }, + change continuous_on ((continuous_multilinear_curry_right_equiv' 𝕜 m E F).symm + ∘ (λ (y : E), p y m.succ)) s, + rw linear_isometry_equiv.comp_continuous_on_iff, + exact H.cont _ A } }, + { rintros ⟨Hzero_eq, Hfderiv_zero, Htaylor⟩, + split, + { exact Hzero_eq }, + { assume m (hm : (m : ℕ∞) < n.succ) x (hx : x ∈ s), + cases m, + { exact Hfderiv_zero x hx }, + { have A : (m : ℕ∞) < n, + by { rw with_top.coe_lt_coe at hm ⊢, exact nat.lt_of_succ_lt_succ hm }, + have : has_fderiv_within_at ((continuous_multilinear_curry_right_equiv' 𝕜 m E F).symm + ∘ (λ (y : E), p y m.succ)) ((p x).shift m.succ).curry_left s x := + Htaylor.fderiv_within _ A x hx, + rw linear_isometry_equiv.comp_has_fderiv_within_at_iff' at this, + convert this, + ext y v, + change (p x (nat.succ (nat.succ m))) (cons y v) + = (p x m.succ.succ) (snoc (cons y (init v)) (v (last _))), + rw [← cons_snoc_eq_snoc_cons, snoc_init_self] } }, + { assume m (hm : (m : ℕ∞) ≤ n.succ), + cases m, + { have : differentiable_on 𝕜 (λ x, p x 0) s := + λ x hx, (Hfderiv_zero x hx).differentiable_within_at, + exact this.continuous_on }, + { have A : (m : ℕ∞) ≤ n, + by { rw with_top.coe_le_coe at hm ⊢, exact nat.lt_succ_iff.mp hm }, + have : continuous_on ((continuous_multilinear_curry_right_equiv' 𝕜 m E F).symm + ∘ (λ (y : E), p y m.succ)) s := + Htaylor.cont _ A, + rwa linear_isometry_equiv.comp_continuous_on_iff at this } } } +end + +/-! ### Smooth functions within a set around a point -/ + +variable (𝕜) + +/-- A function is continuously differentiable up to order `n` within a set `s` at a point `x` if +it admits continuous derivatives up to order `n` in a neighborhood of `x` in `s ∪ {x}`. +For `n = ∞`, we only require that this holds up to any finite order (where the neighborhood may +depend on the finite order we consider). + +For instance, a real function which is `C^m` on `(-1/m, 1/m)` for each natural `m`, but not +better, is `C^∞` at `0` within `univ`. +-/ +def cont_diff_within_at (n : ℕ∞) (f : E → F) (s : set E) (x : E) : Prop := +∀ (m : ℕ), (m : ℕ∞) ≤ n → + ∃ u ∈ 𝓝[insert x s] x, ∃ p : E → formal_multilinear_series 𝕜 E F, + has_ftaylor_series_up_to_on m f p u + +variable {𝕜} + +lemma cont_diff_within_at_nat {n : ℕ} : + cont_diff_within_at 𝕜 n f s x ↔ + ∃ u ∈ 𝓝[insert x s] x, ∃ p : E → formal_multilinear_series 𝕜 E F, + has_ftaylor_series_up_to_on n f p u := +⟨λ H, H n le_rfl, λ ⟨u, hu, p, hp⟩ m hm, ⟨u, hu, p, hp.of_le hm⟩⟩ + +lemma cont_diff_within_at.of_le + (h : cont_diff_within_at 𝕜 n f s x) (hmn : m ≤ n) : + cont_diff_within_at 𝕜 m f s x := +λ k hk, h k (le_trans hk hmn) + +lemma cont_diff_within_at_iff_forall_nat_le : + cont_diff_within_at 𝕜 n f s x ↔ ∀ m : ℕ, ↑m ≤ n → cont_diff_within_at 𝕜 m f s x := +⟨λ H m hm, H.of_le hm, λ H m hm, H m hm _ le_rfl⟩ + +lemma cont_diff_within_at_top : + cont_diff_within_at 𝕜 ∞ f s x ↔ ∀ (n : ℕ), cont_diff_within_at 𝕜 n f s x := +cont_diff_within_at_iff_forall_nat_le.trans $ by simp only [forall_prop_of_true, le_top] + +lemma cont_diff_within_at.continuous_within_at + (h : cont_diff_within_at 𝕜 n f s x) : continuous_within_at f s x := +begin + rcases h 0 bot_le with ⟨u, hu, p, H⟩, + rw [mem_nhds_within_insert] at hu, + exact (H.continuous_on.continuous_within_at hu.1).mono_of_mem hu.2 +end + +lemma cont_diff_within_at.congr_of_eventually_eq + (h : cont_diff_within_at 𝕜 n f s x) (h₁ : f₁ =ᶠ[𝓝[s] x] f) (hx : f₁ x = f x) : + cont_diff_within_at 𝕜 n f₁ s x := +λ m hm, let ⟨u, hu, p, H⟩ := h m hm in +⟨{x ∈ u | f₁ x = f x}, filter.inter_mem hu (mem_nhds_within_insert.2 ⟨hx, h₁⟩), p, + (H.mono (sep_subset _ _)).congr (λ _, and.right)⟩ + +lemma cont_diff_within_at.congr_of_eventually_eq_insert + (h : cont_diff_within_at 𝕜 n f s x) (h₁ : f₁ =ᶠ[𝓝[insert x s] x] f) : + cont_diff_within_at 𝕜 n f₁ s x := +h.congr_of_eventually_eq (nhds_within_mono x (subset_insert x s) h₁) + (mem_of_mem_nhds_within (mem_insert x s) h₁ : _) + +lemma cont_diff_within_at.congr_of_eventually_eq' + (h : cont_diff_within_at 𝕜 n f s x) (h₁ : f₁ =ᶠ[𝓝[s] x] f) (hx : x ∈ s) : + cont_diff_within_at 𝕜 n f₁ s x := +h.congr_of_eventually_eq h₁ $ h₁.self_of_nhds_within hx + +lemma filter.eventually_eq.cont_diff_within_at_iff + (h₁ : f₁ =ᶠ[𝓝[s] x] f) (hx : f₁ x = f x) : + cont_diff_within_at 𝕜 n f₁ s x ↔ cont_diff_within_at 𝕜 n f s x := +⟨λ H, cont_diff_within_at.congr_of_eventually_eq H h₁.symm hx.symm, +λ H, H.congr_of_eventually_eq h₁ hx⟩ + +lemma cont_diff_within_at.congr + (h : cont_diff_within_at 𝕜 n f s x) (h₁ : ∀ y ∈ s, f₁ y = f y) (hx : f₁ x = f x) : + cont_diff_within_at 𝕜 n f₁ s x := +h.congr_of_eventually_eq (filter.eventually_eq_of_mem self_mem_nhds_within h₁) hx + +lemma cont_diff_within_at.congr' + (h : cont_diff_within_at 𝕜 n f s x) (h₁ : ∀ y ∈ s, f₁ y = f y) (hx : x ∈ s) : + cont_diff_within_at 𝕜 n f₁ s x := +h.congr h₁ (h₁ _ hx) + +lemma cont_diff_within_at.mono_of_mem + (h : cont_diff_within_at 𝕜 n f s x) {t : set E} (hst : s ∈ 𝓝[t] x) : + cont_diff_within_at 𝕜 n f t x := +begin + assume m hm, + rcases h m hm with ⟨u, hu, p, H⟩, + exact ⟨u, nhds_within_le_of_mem (insert_mem_nhds_within_insert hst) hu, p, H⟩ +end + +lemma cont_diff_within_at.mono + (h : cont_diff_within_at 𝕜 n f s x) {t : set E} (hst : t ⊆ s) : + cont_diff_within_at 𝕜 n f t x := +h.mono_of_mem $ filter.mem_of_superset self_mem_nhds_within hst + +lemma cont_diff_within_at.congr_nhds + (h : cont_diff_within_at 𝕜 n f s x) {t : set E} (hst : 𝓝[s] x = 𝓝[t] x) : + cont_diff_within_at 𝕜 n f t x := +h.mono_of_mem $ hst ▸ self_mem_nhds_within + +lemma cont_diff_within_at_congr_nhds {t : set E} (hst : 𝓝[s] x = 𝓝[t] x) : + cont_diff_within_at 𝕜 n f s x ↔ cont_diff_within_at 𝕜 n f t x := +⟨λ h, h.congr_nhds hst, λ h, h.congr_nhds hst.symm⟩ + +lemma cont_diff_within_at_inter' (h : t ∈ 𝓝[s] x) : + cont_diff_within_at 𝕜 n f (s ∩ t) x ↔ cont_diff_within_at 𝕜 n f s x := +cont_diff_within_at_congr_nhds $ eq.symm $ nhds_within_restrict'' _ h + +lemma cont_diff_within_at_inter (h : t ∈ 𝓝 x) : + cont_diff_within_at 𝕜 n f (s ∩ t) x ↔ cont_diff_within_at 𝕜 n f s x := +cont_diff_within_at_inter' (mem_nhds_within_of_mem_nhds h) + +lemma cont_diff_within_at_insert {y : E} : + cont_diff_within_at 𝕜 n f (insert y s) x ↔ cont_diff_within_at 𝕜 n f s x := +begin + simp_rw [cont_diff_within_at], + rcases eq_or_ne x y with rfl|h, + { simp_rw [insert_eq_of_mem (mem_insert _ _)] }, + simp_rw [insert_comm x y, nhds_within_insert_of_ne h] +end + +alias cont_diff_within_at_insert ↔ cont_diff_within_at.of_insert cont_diff_within_at.insert' + +lemma cont_diff_within_at.insert (h : cont_diff_within_at 𝕜 n f s x) : + cont_diff_within_at 𝕜 n f (insert x s) x := +h.insert' + +/-- If a function is `C^n` within a set at a point, with `n ≥ 1`, then it is differentiable +within this set at this point. -/ +lemma cont_diff_within_at.differentiable_within_at' + (h : cont_diff_within_at 𝕜 n f s x) (hn : 1 ≤ n) : + differentiable_within_at 𝕜 f (insert x s) x := +begin + rcases h 1 hn with ⟨u, hu, p, H⟩, + rcases mem_nhds_within.1 hu with ⟨t, t_open, xt, tu⟩, + rw inter_comm at tu, + have := ((H.mono tu).differentiable_on le_rfl) x ⟨mem_insert x s, xt⟩, + exact (differentiable_within_at_inter (is_open.mem_nhds t_open xt)).1 this, +end + +lemma cont_diff_within_at.differentiable_within_at + (h : cont_diff_within_at 𝕜 n f s x) (hn : 1 ≤ n) : + differentiable_within_at 𝕜 f s x := +(h.differentiable_within_at' hn).mono (subset_insert x s) + +/-- A function is `C^(n + 1)` on a domain iff locally, it has a derivative which is `C^n`. -/ +theorem cont_diff_within_at_succ_iff_has_fderiv_within_at {n : ℕ} : + cont_diff_within_at 𝕜 ((n + 1) : ℕ) f s x + ↔ ∃ u ∈ 𝓝[insert x s] x, ∃ f' : E → (E →L[𝕜] F), + (∀ x ∈ u, has_fderiv_within_at f (f' x) u x) ∧ (cont_diff_within_at 𝕜 n f' u x) := +begin + split, + { assume h, + rcases h n.succ le_rfl with ⟨u, hu, p, Hp⟩, + refine ⟨u, hu, λ y, (continuous_multilinear_curry_fin1 𝕜 E F) (p y 1), + λ y hy, Hp.has_fderiv_within_at (with_top.coe_le_coe.2 (nat.le_add_left 1 n)) hy, _⟩, + assume m hm, + refine ⟨u, _, λ (y : E), (p y).shift, _⟩, + { convert self_mem_nhds_within, + have : x ∈ insert x s, by simp, + exact (insert_eq_of_mem (mem_of_mem_nhds_within this hu)) }, + { rw has_ftaylor_series_up_to_on_succ_iff_right at Hp, + exact Hp.2.2.of_le hm } }, + { rintros ⟨u, hu, f', f'_eq_deriv, Hf'⟩, + rw cont_diff_within_at_nat, + rcases Hf' n le_rfl with ⟨v, hv, p', Hp'⟩, + refine ⟨v ∩ u, _, λ x, (p' x).unshift (f x), _⟩, + { apply filter.inter_mem _ hu, + apply nhds_within_le_of_mem hu, + exact nhds_within_mono _ (subset_insert x u) hv }, + { rw has_ftaylor_series_up_to_on_succ_iff_right, + refine ⟨λ y hy, rfl, λ y hy, _, _⟩, + { change has_fderiv_within_at (λ z, (continuous_multilinear_curry_fin0 𝕜 E F).symm (f z)) + ((formal_multilinear_series.unshift (p' y) (f y) 1).curry_left) (v ∩ u) y, + rw linear_isometry_equiv.comp_has_fderiv_within_at_iff', + convert (f'_eq_deriv y hy.2).mono (inter_subset_right v u), + rw ← Hp'.zero_eq y hy.1, + ext z, + change ((p' y 0) (init (@cons 0 (λ i, E) z 0))) (@cons 0 (λ i, E) z 0 (last 0)) + = ((p' y 0) 0) z, + unfold_coes, + congr }, + { convert (Hp'.mono (inter_subset_left v u)).congr (λ x hx, Hp'.zero_eq x hx.1), + { ext x y, + change p' x 0 (init (@snoc 0 (λ i : fin 1, E) 0 y)) y = p' x 0 0 y, + rw init_snoc }, + { ext x k v y, + change p' x k (init (@snoc k (λ i : fin k.succ, E) v y)) + (@snoc k (λ i : fin k.succ, E) v y (last k)) = p' x k v y, + rw [snoc_last, init_snoc] } } } } +end + +/-- A version of `cont_diff_within_at_succ_iff_has_fderiv_within_at` where all derivatives + are taken within the same set. -/ +lemma cont_diff_within_at_succ_iff_has_fderiv_within_at' {n : ℕ} : + cont_diff_within_at 𝕜 (n + 1 : ℕ) f s x + ↔ ∃ u ∈ 𝓝[insert x s] x, u ⊆ insert x s ∧ ∃ f' : E → E →L[𝕜] F, + (∀ x ∈ u, has_fderiv_within_at f (f' x) s x) ∧ cont_diff_within_at 𝕜 n f' s x := +begin + refine ⟨λ hf, _, _⟩, + { obtain ⟨u, hu, f', huf', hf'⟩ := cont_diff_within_at_succ_iff_has_fderiv_within_at.mp hf, + obtain ⟨w, hw, hxw, hwu⟩ := mem_nhds_within.mp hu, + rw [inter_comm] at hwu, + refine ⟨insert x s ∩ w, inter_mem_nhds_within _ (hw.mem_nhds hxw), inter_subset_left _ _, + f', λ y hy, _, _⟩, + { refine ((huf' y $ hwu hy).mono hwu).mono_of_mem _, + refine mem_of_superset _ (inter_subset_inter_left _ (subset_insert _ _)), + refine inter_mem_nhds_within _ (hw.mem_nhds hy.2) }, + { exact hf'.mono_of_mem (nhds_within_mono _ (subset_insert _ _) hu) } }, + { rw [← cont_diff_within_at_insert, cont_diff_within_at_succ_iff_has_fderiv_within_at, + insert_eq_of_mem (mem_insert _ _)], + rintro ⟨u, hu, hus, f', huf', hf'⟩, + refine ⟨u, hu, f', λ y hy, (huf' y hy).insert'.mono hus, hf'.insert.mono hus⟩ } +end + +/-! ### Smooth functions within a set -/ + +variable (𝕜) + +/-- A function is continuously differentiable up to `n` on `s` if, for any point `x` in `s`, it +admits continuous derivatives up to order `n` on a neighborhood of `x` in `s`. + +For `n = ∞`, we only require that this holds up to any finite order (where the neighborhood may +depend on the finite order we consider). +-/ +def cont_diff_on (n : ℕ∞) (f : E → F) (s : set E) : Prop := +∀ x ∈ s, cont_diff_within_at 𝕜 n f s x + +variable {𝕜} + +lemma cont_diff_on.cont_diff_within_at (h : cont_diff_on 𝕜 n f s) (hx : x ∈ s) : + cont_diff_within_at 𝕜 n f s x := +h x hx + +lemma cont_diff_within_at.cont_diff_on {m : ℕ} + (hm : (m : ℕ∞) ≤ n) (h : cont_diff_within_at 𝕜 n f s x) : + ∃ u ∈ 𝓝[insert x s] x, u ⊆ insert x s ∧ cont_diff_on 𝕜 m f u := +begin + rcases h m hm with ⟨u, u_nhd, p, hp⟩, + refine ⟨u ∩ insert x s, filter.inter_mem u_nhd self_mem_nhds_within, + inter_subset_right _ _, _⟩, + assume y hy m' hm', + refine ⟨u ∩ insert x s, _, p, (hp.mono (inter_subset_left _ _)).of_le hm'⟩, + convert self_mem_nhds_within, + exact insert_eq_of_mem hy +end + +protected lemma cont_diff_within_at.eventually {n : ℕ} + (h : cont_diff_within_at 𝕜 n f s x) : + ∀ᶠ y in 𝓝[insert x s] x, cont_diff_within_at 𝕜 n f s y := +begin + rcases h.cont_diff_on le_rfl with ⟨u, hu, hu_sub, hd⟩, + have : ∀ᶠ (y : E) in 𝓝[insert x s] x, u ∈ 𝓝[insert x s] y ∧ y ∈ u, + from (eventually_nhds_within_nhds_within.2 hu).and hu, + refine this.mono (λ y hy, (hd y hy.2).mono_of_mem _), + exact nhds_within_mono y (subset_insert _ _) hy.1 +end + +lemma cont_diff_on.of_le (h : cont_diff_on 𝕜 n f s) (hmn : m ≤ n) : + cont_diff_on 𝕜 m f s := +λ x hx, (h x hx).of_le hmn + +lemma cont_diff_on.of_succ {n : ℕ} (h : cont_diff_on 𝕜 (n + 1) f s) : cont_diff_on 𝕜 n f s := +h.of_le $ with_top.coe_le_coe.mpr le_self_add + +lemma cont_diff_on.one_of_succ {n : ℕ} (h : cont_diff_on 𝕜 (n + 1) f s) : cont_diff_on 𝕜 1 f s := +h.of_le $ with_top.coe_le_coe.mpr le_add_self + +lemma cont_diff_on_iff_forall_nat_le : + cont_diff_on 𝕜 n f s ↔ ∀ m : ℕ, ↑m ≤ n → cont_diff_on 𝕜 m f s := +⟨λ H m hm, H.of_le hm, λ H x hx m hm, H m hm x hx m le_rfl⟩ + +lemma cont_diff_on_top : + cont_diff_on 𝕜 ∞ f s ↔ ∀ (n : ℕ), cont_diff_on 𝕜 n f s := +cont_diff_on_iff_forall_nat_le.trans $ by simp only [le_top, forall_prop_of_true] + +lemma cont_diff_on_all_iff_nat : + (∀ n, cont_diff_on 𝕜 n f s) ↔ (∀ n : ℕ, cont_diff_on 𝕜 n f s) := +begin + refine ⟨λ H n, H n, _⟩, + rintro H (_|n), + exacts [cont_diff_on_top.2 H, H n] +end + +lemma cont_diff_on.continuous_on + (h : cont_diff_on 𝕜 n f s) : continuous_on f s := +λ x hx, (h x hx).continuous_within_at + +lemma cont_diff_on.congr + (h : cont_diff_on 𝕜 n f s) (h₁ : ∀ x ∈ s, f₁ x = f x) : + cont_diff_on 𝕜 n f₁ s := +λ x hx, (h x hx).congr h₁ (h₁ x hx) + +lemma cont_diff_on_congr (h₁ : ∀ x ∈ s, f₁ x = f x) : + cont_diff_on 𝕜 n f₁ s ↔ cont_diff_on 𝕜 n f s := +⟨λ H, H.congr (λ x hx, (h₁ x hx).symm), λ H, H.congr h₁⟩ + +lemma cont_diff_on.mono + (h : cont_diff_on 𝕜 n f s) {t : set E} (hst : t ⊆ s) : + cont_diff_on 𝕜 n f t := +λ x hx, (h x (hst hx)).mono hst + +lemma cont_diff_on.congr_mono + (hf : cont_diff_on 𝕜 n f s) (h₁ : ∀ x ∈ s₁, f₁ x = f x) (hs : s₁ ⊆ s) : + cont_diff_on 𝕜 n f₁ s₁ := +(hf.mono hs).congr h₁ + +/-- If a function is `C^n` on a set with `n ≥ 1`, then it is differentiable there. -/ +lemma cont_diff_on.differentiable_on + (h : cont_diff_on 𝕜 n f s) (hn : 1 ≤ n) : differentiable_on 𝕜 f s := +λ x hx, (h x hx).differentiable_within_at hn + +/-- If a function is `C^n` around each point in a set, then it is `C^n` on the set. -/ +lemma cont_diff_on_of_locally_cont_diff_on + (h : ∀ x ∈ s, ∃u, is_open u ∧ x ∈ u ∧ cont_diff_on 𝕜 n f (s ∩ u)) : + cont_diff_on 𝕜 n f s := +begin + assume x xs, + rcases h x xs with ⟨u, u_open, xu, hu⟩, + apply (cont_diff_within_at_inter _).1 (hu x ⟨xs, xu⟩), + exact is_open.mem_nhds u_open xu +end + +/-- A function is `C^(n + 1)` on a domain iff locally, it has a derivative which is `C^n`. -/ +theorem cont_diff_on_succ_iff_has_fderiv_within_at {n : ℕ} : + cont_diff_on 𝕜 ((n + 1) : ℕ) f s + ↔ ∀ x ∈ s, ∃ u ∈ 𝓝[insert x s] x, ∃ f' : E → (E →L[𝕜] F), + (∀ x ∈ u, has_fderiv_within_at f (f' x) u x) ∧ (cont_diff_on 𝕜 n f' u) := +begin + split, + { assume h x hx, + rcases (h x hx) n.succ le_rfl with ⟨u, hu, p, Hp⟩, + refine ⟨u, hu, λ y, (continuous_multilinear_curry_fin1 𝕜 E F) (p y 1), + λ y hy, Hp.has_fderiv_within_at (with_top.coe_le_coe.2 (nat.le_add_left 1 n)) hy, _⟩, + rw has_ftaylor_series_up_to_on_succ_iff_right at Hp, + assume z hz m hm, + refine ⟨u, _, λ (x : E), (p x).shift, Hp.2.2.of_le hm⟩, + convert self_mem_nhds_within, + exact insert_eq_of_mem hz, }, + { assume h x hx, + rw cont_diff_within_at_succ_iff_has_fderiv_within_at, + rcases h x hx with ⟨u, u_nhbd, f', hu, hf'⟩, + have : x ∈ u := mem_of_mem_nhds_within (mem_insert _ _) u_nhbd, + exact ⟨u, u_nhbd, f', hu, hf' x this⟩ } +end + +/-! ### Iterated derivative within a set -/ +variable (𝕜) + +/-- +The `n`-th derivative of a function along a set, defined inductively by saying that the `n+1`-th +derivative of `f` is the derivative of the `n`-th derivative of `f` along this set, together with +an uncurrying step to see it as a multilinear map in `n+1` variables.. +-/ +noncomputable def iterated_fderiv_within (n : ℕ) (f : E → F) (s : set E) : + E → (E [×n]→L[𝕜] F) := +nat.rec_on n + (λ x, continuous_multilinear_map.curry0 𝕜 E (f x)) + (λ n rec x, continuous_linear_map.uncurry_left (fderiv_within 𝕜 rec s x)) + +/-- Formal Taylor series associated to a function within a set. -/ +def ftaylor_series_within (f : E → F) (s : set E) (x : E) : formal_multilinear_series 𝕜 E F := +λ n, iterated_fderiv_within 𝕜 n f s x + +variable {𝕜} + +@[simp] lemma iterated_fderiv_within_zero_apply (m : (fin 0) → E) : + (iterated_fderiv_within 𝕜 0 f s x : ((fin 0) → E) → F) m = f x := rfl + +lemma iterated_fderiv_within_zero_eq_comp : + iterated_fderiv_within 𝕜 0 f s = (continuous_multilinear_curry_fin0 𝕜 E F).symm ∘ f := rfl + +@[simp] lemma norm_iterated_fderiv_within_zero : + ‖iterated_fderiv_within 𝕜 0 f s x‖ = ‖f x‖ := +by rw [iterated_fderiv_within_zero_eq_comp, linear_isometry_equiv.norm_map] + +lemma iterated_fderiv_within_succ_apply_left {n : ℕ} (m : fin (n + 1) → E): + (iterated_fderiv_within 𝕜 (n + 1) f s x : (fin (n + 1) → E) → F) m + = (fderiv_within 𝕜 (iterated_fderiv_within 𝕜 n f s) s x : E → (E [×n]→L[𝕜] F)) + (m 0) (tail m) := rfl + +/-- Writing explicitly the `n+1`-th derivative as the composition of a currying linear equiv, +and the derivative of the `n`-th derivative. -/ +lemma iterated_fderiv_within_succ_eq_comp_left {n : ℕ} : + iterated_fderiv_within 𝕜 (n + 1) f s = + (continuous_multilinear_curry_left_equiv 𝕜 (λ(i : fin (n + 1)), E) F) + ∘ (fderiv_within 𝕜 (iterated_fderiv_within 𝕜 n f s) s) := rfl + +lemma norm_fderiv_within_iterated_fderiv_within {n : ℕ} : + ‖fderiv_within 𝕜 (iterated_fderiv_within 𝕜 n f s) s x‖ = + ‖iterated_fderiv_within 𝕜 (n + 1) f s x‖ := +by rw [iterated_fderiv_within_succ_eq_comp_left, linear_isometry_equiv.norm_map] + +theorem iterated_fderiv_within_succ_apply_right {n : ℕ} + (hs : unique_diff_on 𝕜 s) (hx : x ∈ s) (m : fin (n + 1) → E) : + (iterated_fderiv_within 𝕜 (n + 1) f s x : (fin (n + 1) → E) → F) m + = iterated_fderiv_within 𝕜 n (λy, fderiv_within 𝕜 f s y) s x (init m) (m (last n)) := +begin + induction n with n IH generalizing x, + { rw [iterated_fderiv_within_succ_eq_comp_left, iterated_fderiv_within_zero_eq_comp, + iterated_fderiv_within_zero_apply, + function.comp_apply, linear_isometry_equiv.comp_fderiv_within _ (hs x hx)], + refl }, + { let I := continuous_multilinear_curry_right_equiv' 𝕜 n E F, + have A : ∀ y ∈ s, iterated_fderiv_within 𝕜 n.succ f s y + = (I ∘ (iterated_fderiv_within 𝕜 n (λy, fderiv_within 𝕜 f s y) s)) y, + by { assume y hy, ext m, rw @IH m y hy, refl }, + calc + (iterated_fderiv_within 𝕜 (n+2) f s x : (fin (n+2) → E) → F) m = + (fderiv_within 𝕜 (iterated_fderiv_within 𝕜 n.succ f s) s x + : E → (E [×(n + 1)]→L[𝕜] F)) (m 0) (tail m) : rfl + ... = (fderiv_within 𝕜 (I ∘ (iterated_fderiv_within 𝕜 n (fderiv_within 𝕜 f s) s)) s x + : E → (E [×(n + 1)]→L[𝕜] F)) (m 0) (tail m) : + by rw fderiv_within_congr (hs x hx) A (A x hx) + ... = (I ∘ fderiv_within 𝕜 ((iterated_fderiv_within 𝕜 n (fderiv_within 𝕜 f s) s)) s x + : E → (E [×(n + 1)]→L[𝕜] F)) (m 0) (tail m) : + by { rw linear_isometry_equiv.comp_fderiv_within _ (hs x hx), refl } + ... = (fderiv_within 𝕜 ((iterated_fderiv_within 𝕜 n (λ y, fderiv_within 𝕜 f s y) s)) s x + : E → (E [×n]→L[𝕜] (E →L[𝕜] F))) (m 0) (init (tail m)) ((tail m) (last n)) : rfl + ... = iterated_fderiv_within 𝕜 (nat.succ n) (λ y, fderiv_within 𝕜 f s y) s x + (init m) (m (last (n + 1))) : + by { rw [iterated_fderiv_within_succ_apply_left, tail_init_eq_init_tail], refl } } +end + +/-- Writing explicitly the `n+1`-th derivative as the composition of a currying linear equiv, +and the `n`-th derivative of the derivative. -/ +lemma iterated_fderiv_within_succ_eq_comp_right {n : ℕ} (hs : unique_diff_on 𝕜 s) (hx : x ∈ s) : + iterated_fderiv_within 𝕜 (n + 1) f s x = + ((continuous_multilinear_curry_right_equiv' 𝕜 n E F) + ∘ (iterated_fderiv_within 𝕜 n (λy, fderiv_within 𝕜 f s y) s)) x := +by { ext m, rw iterated_fderiv_within_succ_apply_right hs hx, refl } + +lemma norm_iterated_fderiv_within_fderiv_within {n : ℕ} (hs : unique_diff_on 𝕜 s) (hx : x ∈ s) : + ‖iterated_fderiv_within 𝕜 n (fderiv_within 𝕜 f s) s x‖ = + ‖iterated_fderiv_within 𝕜 (n + 1) f s x‖ := +by rw [iterated_fderiv_within_succ_eq_comp_right hs hx, linear_isometry_equiv.norm_map] + +@[simp] lemma iterated_fderiv_within_one_apply + (hs : unique_diff_on 𝕜 s) (hx : x ∈ s) (m : (fin 1) → E) : + (iterated_fderiv_within 𝕜 1 f s x : ((fin 1) → E) → F) m + = (fderiv_within 𝕜 f s x : E → F) (m 0) := +by { rw [iterated_fderiv_within_succ_apply_right hs hx, iterated_fderiv_within_zero_apply], refl } + +/-- If two functions coincide on a set `s` of unique differentiability, then their iterated +differentials within this set coincide. -/ +lemma iterated_fderiv_within_congr {n : ℕ} + (hs : unique_diff_on 𝕜 s) (hL : ∀y∈s, f₁ y = f y) (hx : x ∈ s) : + iterated_fderiv_within 𝕜 n f₁ s x = iterated_fderiv_within 𝕜 n f s x := +begin + induction n with n IH generalizing x, + { ext m, simp [hL x hx] }, + { have : fderiv_within 𝕜 (λ y, iterated_fderiv_within 𝕜 n f₁ s y) s x + = fderiv_within 𝕜 (λ y, iterated_fderiv_within 𝕜 n f s y) s x := + fderiv_within_congr (hs x hx) (λ y hy, IH hy) (IH hx), + ext m, + rw [iterated_fderiv_within_succ_apply_left, iterated_fderiv_within_succ_apply_left, this] } +end + +/-- The iterated differential within a set `s` at a point `x` is not modified if one intersects +`s` with an open set containing `x`. -/ +lemma iterated_fderiv_within_inter_open {n : ℕ} (hu : is_open u) + (hs : unique_diff_on 𝕜 (s ∩ u)) (hx : x ∈ s ∩ u) : + iterated_fderiv_within 𝕜 n f (s ∩ u) x = iterated_fderiv_within 𝕜 n f s x := +begin + induction n with n IH generalizing x, + { ext m, simp }, + { have A : fderiv_within 𝕜 (λ y, iterated_fderiv_within 𝕜 n f (s ∩ u) y) (s ∩ u) x + = fderiv_within 𝕜 (λ y, iterated_fderiv_within 𝕜 n f s y) (s ∩ u) x := + fderiv_within_congr (hs x hx) (λ y hy, IH hy) (IH hx), + have B : fderiv_within 𝕜 (λ y, iterated_fderiv_within 𝕜 n f s y) (s ∩ u) x + = fderiv_within 𝕜 (λ y, iterated_fderiv_within 𝕜 n f s y) s x := + fderiv_within_inter (is_open.mem_nhds hu hx.2) + ((unique_diff_within_at_inter (is_open.mem_nhds hu hx.2)).1 (hs x hx)), + ext m, + rw [iterated_fderiv_within_succ_apply_left, iterated_fderiv_within_succ_apply_left, A, B] } +end + +/-- The iterated differential within a set `s` at a point `x` is not modified if one intersects +`s` with a neighborhood of `x` within `s`. -/ +lemma iterated_fderiv_within_inter' {n : ℕ} + (hu : u ∈ 𝓝[s] x) (hs : unique_diff_on 𝕜 s) (xs : x ∈ s) : + iterated_fderiv_within 𝕜 n f (s ∩ u) x = iterated_fderiv_within 𝕜 n f s x := +begin + obtain ⟨v, v_open, xv, vu⟩ : ∃ v, is_open v ∧ x ∈ v ∧ v ∩ s ⊆ u := mem_nhds_within.1 hu, + have A : (s ∩ u) ∩ v = s ∩ v, + { apply subset.antisymm (inter_subset_inter (inter_subset_left _ _) (subset.refl _)), + exact λ y ⟨ys, yv⟩, ⟨⟨ys, vu ⟨yv, ys⟩⟩, yv⟩ }, + have : iterated_fderiv_within 𝕜 n f (s ∩ v) x = iterated_fderiv_within 𝕜 n f s x := + iterated_fderiv_within_inter_open v_open (hs.inter v_open) ⟨xs, xv⟩, + rw ← this, + have : iterated_fderiv_within 𝕜 n f ((s ∩ u) ∩ v) x = iterated_fderiv_within 𝕜 n f (s ∩ u) x, + { refine iterated_fderiv_within_inter_open v_open _ ⟨⟨xs, vu ⟨xv, xs⟩⟩, xv⟩, + rw A, + exact hs.inter v_open }, + rw A at this, + rw ← this +end + +/-- The iterated differential within a set `s` at a point `x` is not modified if one intersects +`s` with a neighborhood of `x`. -/ +lemma iterated_fderiv_within_inter {n : ℕ} + (hu : u ∈ 𝓝 x) (hs : unique_diff_on 𝕜 s) (xs : x ∈ s) : + iterated_fderiv_within 𝕜 n f (s ∩ u) x = iterated_fderiv_within 𝕜 n f s x := +iterated_fderiv_within_inter' (mem_nhds_within_of_mem_nhds hu) hs xs + +@[simp] lemma cont_diff_on_zero : + cont_diff_on 𝕜 0 f s ↔ continuous_on f s := +begin + refine ⟨λ H, H.continuous_on, λ H, _⟩, + assume x hx m hm, + have : (m : ℕ∞) = 0 := le_antisymm hm bot_le, + rw this, + refine ⟨insert x s, self_mem_nhds_within, ftaylor_series_within 𝕜 f s, _⟩, + rw has_ftaylor_series_up_to_on_zero_iff, + exact ⟨by rwa insert_eq_of_mem hx, λ x hx, by simp [ftaylor_series_within]⟩ +end + +lemma cont_diff_within_at_zero (hx : x ∈ s) : + cont_diff_within_at 𝕜 0 f s x ↔ ∃ u ∈ 𝓝[s] x, continuous_on f (s ∩ u) := +begin + split, + { intros h, + obtain ⟨u, H, p, hp⟩ := h 0 (by norm_num), + refine ⟨u, _, _⟩, + { simpa [hx] using H }, + { simp only [with_top.coe_zero, has_ftaylor_series_up_to_on_zero_iff] at hp, + exact hp.1.mono (inter_subset_right s u) } }, + { rintros ⟨u, H, hu⟩, + rw ← cont_diff_within_at_inter' H, + have h' : x ∈ s ∩ u := ⟨hx, mem_of_mem_nhds_within hx H⟩, + exact (cont_diff_on_zero.mpr hu).cont_diff_within_at h' } +end + +/-- On a set with unique differentiability, any choice of iterated differential has to coincide +with the one we have chosen in `iterated_fderiv_within 𝕜 m f s`. -/ +theorem has_ftaylor_series_up_to_on.eq_ftaylor_series_of_unique_diff_on + (h : has_ftaylor_series_up_to_on n f p s) + {m : ℕ} (hmn : (m : ℕ∞) ≤ n) (hs : unique_diff_on 𝕜 s) (hx : x ∈ s) : + p x m = iterated_fderiv_within 𝕜 m f s x := +begin + induction m with m IH generalizing x, + { rw [h.zero_eq' hx, iterated_fderiv_within_zero_eq_comp] }, + { have A : (m : ℕ∞) < n := lt_of_lt_of_le (with_top.coe_lt_coe.2 (lt_add_one m)) hmn, + have : has_fderiv_within_at (λ (y : E), iterated_fderiv_within 𝕜 m f s y) + (continuous_multilinear_map.curry_left (p x (nat.succ m))) s x := + (h.fderiv_within m A x hx).congr (λ y hy, (IH (le_of_lt A) hy).symm) (IH (le_of_lt A) hx).symm, + rw [iterated_fderiv_within_succ_eq_comp_left, function.comp_apply, + this.fderiv_within (hs x hx)], + exact (continuous_multilinear_map.uncurry_curry_left _).symm } +end + +/-- When a function is `C^n` in a set `s` of unique differentiability, it admits +`ftaylor_series_within 𝕜 f s` as a Taylor series up to order `n` in `s`. -/ +theorem cont_diff_on.ftaylor_series_within + (h : cont_diff_on 𝕜 n f s) (hs : unique_diff_on 𝕜 s) : + has_ftaylor_series_up_to_on n f (ftaylor_series_within 𝕜 f s) s := +begin + split, + { assume x hx, + simp only [ftaylor_series_within, continuous_multilinear_map.uncurry0_apply, + iterated_fderiv_within_zero_apply] }, + { assume m hm x hx, + rcases (h x hx) m.succ (enat.add_one_le_of_lt hm) with ⟨u, hu, p, Hp⟩, + rw insert_eq_of_mem hx at hu, + rcases mem_nhds_within.1 hu with ⟨o, o_open, xo, ho⟩, + rw inter_comm at ho, + have : p x m.succ = ftaylor_series_within 𝕜 f s x m.succ, + { change p x m.succ = iterated_fderiv_within 𝕜 m.succ f s x, + rw ← iterated_fderiv_within_inter (is_open.mem_nhds o_open xo) hs hx, + exact (Hp.mono ho).eq_ftaylor_series_of_unique_diff_on le_rfl + (hs.inter o_open) ⟨hx, xo⟩ }, + rw [← this, ← has_fderiv_within_at_inter (is_open.mem_nhds o_open xo)], + have A : ∀ y ∈ s ∩ o, p y m = ftaylor_series_within 𝕜 f s y m, + { rintros y ⟨hy, yo⟩, + change p y m = iterated_fderiv_within 𝕜 m f s y, + rw ← iterated_fderiv_within_inter (is_open.mem_nhds o_open yo) hs hy, + exact (Hp.mono ho).eq_ftaylor_series_of_unique_diff_on (with_top.coe_le_coe.2 (nat.le_succ m)) + (hs.inter o_open) ⟨hy, yo⟩ }, + exact ((Hp.mono ho).fderiv_within m (with_top.coe_lt_coe.2 (lt_add_one m)) x ⟨hx, xo⟩).congr + (λ y hy, (A y hy).symm) (A x ⟨hx, xo⟩).symm }, + { assume m hm, + apply continuous_on_of_locally_continuous_on, + assume x hx, + rcases h x hx m hm with ⟨u, hu, p, Hp⟩, + rcases mem_nhds_within.1 hu with ⟨o, o_open, xo, ho⟩, + rw insert_eq_of_mem hx at ho, + rw inter_comm at ho, + refine ⟨o, o_open, xo, _⟩, + have A : ∀ y ∈ s ∩ o, p y m = ftaylor_series_within 𝕜 f s y m, + { rintros y ⟨hy, yo⟩, + change p y m = iterated_fderiv_within 𝕜 m f s y, + rw ← iterated_fderiv_within_inter (is_open.mem_nhds o_open yo) hs hy, + exact (Hp.mono ho).eq_ftaylor_series_of_unique_diff_on le_rfl + (hs.inter o_open) ⟨hy, yo⟩ }, + exact ((Hp.mono ho).cont m le_rfl).congr (λ y hy, (A y hy).symm) } +end + +lemma cont_diff_on_of_continuous_on_differentiable_on + (Hcont : ∀ (m : ℕ), (m : ℕ∞) ≤ n → + continuous_on (λ x, iterated_fderiv_within 𝕜 m f s x) s) + (Hdiff : ∀ (m : ℕ), (m : ℕ∞) < n → + differentiable_on 𝕜 (λ x, iterated_fderiv_within 𝕜 m f s x) s) : + cont_diff_on 𝕜 n f s := +begin + assume x hx m hm, + rw insert_eq_of_mem hx, + refine ⟨s, self_mem_nhds_within, ftaylor_series_within 𝕜 f s, _⟩, + split, + { assume y hy, + simp only [ftaylor_series_within, continuous_multilinear_map.uncurry0_apply, + iterated_fderiv_within_zero_apply] }, + { assume k hk y hy, + convert (Hdiff k (lt_of_lt_of_le hk hm) y hy).has_fderiv_within_at, + simp only [ftaylor_series_within, iterated_fderiv_within_succ_eq_comp_left, + continuous_linear_equiv.coe_apply, function.comp_app, coe_fn_coe_base], + exact continuous_linear_map.curry_uncurry_left _ }, + { assume k hk, + exact Hcont k (le_trans hk hm) } +end + +lemma cont_diff_on_of_differentiable_on + (h : ∀(m : ℕ), (m : ℕ∞) ≤ n → differentiable_on 𝕜 (iterated_fderiv_within 𝕜 m f s) s) : + cont_diff_on 𝕜 n f s := +cont_diff_on_of_continuous_on_differentiable_on + (λ m hm, (h m hm).continuous_on) (λ m hm, (h m (le_of_lt hm))) + +lemma cont_diff_on.continuous_on_iterated_fderiv_within {m : ℕ} + (h : cont_diff_on 𝕜 n f s) (hmn : (m : ℕ∞) ≤ n) (hs : unique_diff_on 𝕜 s) : + continuous_on (iterated_fderiv_within 𝕜 m f s) s := +(h.ftaylor_series_within hs).cont m hmn + +lemma cont_diff_on.differentiable_on_iterated_fderiv_within {m : ℕ} + (h : cont_diff_on 𝕜 n f s) (hmn : (m : ℕ∞) < n) (hs : unique_diff_on 𝕜 s) : + differentiable_on 𝕜 (iterated_fderiv_within 𝕜 m f s) s := +λ x hx, ((h.ftaylor_series_within hs).fderiv_within m hmn x hx).differentiable_within_at + +lemma cont_diff_on_iff_continuous_on_differentiable_on + (hs : unique_diff_on 𝕜 s) : + cont_diff_on 𝕜 n f s ↔ + (∀ (m : ℕ), (m : ℕ∞) ≤ n → + continuous_on (λ x, iterated_fderiv_within 𝕜 m f s x) s) + ∧ (∀ (m : ℕ), (m : ℕ∞) < n → + differentiable_on 𝕜 (λ x, iterated_fderiv_within 𝕜 m f s x) s) := +begin + split, + { assume h, + split, + { assume m hm, exact h.continuous_on_iterated_fderiv_within hm hs }, + { assume m hm, exact h.differentiable_on_iterated_fderiv_within hm hs } }, + { assume h, + 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 + 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 +differentiable there, and its derivative (expressed with `fderiv`) is `C^n`. -/ +theorem cont_diff_on_succ_iff_fderiv_of_open {n : ℕ} (hs : is_open s) : + cont_diff_on 𝕜 ((n + 1) : ℕ) f s ↔ + differentiable_on 𝕜 f s ∧ cont_diff_on 𝕜 n (λ y, fderiv 𝕜 f y) s := +begin + rw cont_diff_on_succ_iff_fderiv_within hs.unique_diff_on, + congrm _ ∧ _, + apply cont_diff_on_congr, + assume x hx, + exact fderiv_within_of_open hs hx +end + +/-- A function is `C^∞` on a domain with unique derivatives if and only if it is differentiable +there, and its derivative (expressed with `fderiv_within`) is `C^∞`. -/ +theorem cont_diff_on_top_iff_fderiv_within (hs : unique_diff_on 𝕜 s) : + cont_diff_on 𝕜 ∞ f s ↔ + differentiable_on 𝕜 f s ∧ cont_diff_on 𝕜 ∞ (λ y, fderiv_within 𝕜 f s y) s := +begin + split, + { assume h, + refine ⟨h.differentiable_on le_top, _⟩, + apply cont_diff_on_top.2 (λ n, ((cont_diff_on_succ_iff_fderiv_within hs).1 _).2), + exact h.of_le le_top }, + { assume h, + refine cont_diff_on_top.2 (λ n, _), + have A : (n : ℕ∞) ≤ ∞ := le_top, + apply ((cont_diff_on_succ_iff_fderiv_within hs).2 ⟨h.1, h.2.of_le A⟩).of_le, + exact with_top.coe_le_coe.2 (nat.le_succ n) } +end + +/-- A function is `C^∞` on an open domain if and only if it is differentiable there, and its +derivative (expressed with `fderiv`) is `C^∞`. -/ +theorem cont_diff_on_top_iff_fderiv_of_open (hs : is_open s) : + cont_diff_on 𝕜 ∞ f s ↔ + differentiable_on 𝕜 f s ∧ cont_diff_on 𝕜 ∞ (λ y, fderiv 𝕜 f y) s := +begin + rw cont_diff_on_top_iff_fderiv_within hs.unique_diff_on, + congrm _ ∧ _, + apply cont_diff_on_congr, + assume x hx, + exact fderiv_within_of_open hs hx +end + +lemma cont_diff_on.fderiv_within + (hf : cont_diff_on 𝕜 n f s) (hs : unique_diff_on 𝕜 s) (hmn : m + 1 ≤ n) : + cont_diff_on 𝕜 m (λ y, fderiv_within 𝕜 f s y) s := +begin + cases m, + { change ∞ + 1 ≤ n at hmn, + have : n = ∞, by simpa using hmn, + rw this at hf, + exact ((cont_diff_on_top_iff_fderiv_within hs).1 hf).2 }, + { change (m.succ : ℕ∞) ≤ n at hmn, + exact ((cont_diff_on_succ_iff_fderiv_within hs).1 (hf.of_le hmn)).2 } +end + +lemma cont_diff_on.fderiv_of_open + (hf : cont_diff_on 𝕜 n f s) (hs : is_open s) (hmn : m + 1 ≤ n) : + cont_diff_on 𝕜 m (λ y, fderiv 𝕜 f y) s := +(hf.fderiv_within hs.unique_diff_on hmn).congr (λ x hx, (fderiv_within_of_open hs hx).symm) + +lemma cont_diff_on.continuous_on_fderiv_within + (h : cont_diff_on 𝕜 n f s) (hs : unique_diff_on 𝕜 s) (hn : 1 ≤ n) : + continuous_on (λ x, fderiv_within 𝕜 f s x) s := +((cont_diff_on_succ_iff_fderiv_within hs).1 (h.of_le hn)).2.continuous_on + +lemma cont_diff_on.continuous_on_fderiv_of_open + (h : cont_diff_on 𝕜 n f s) (hs : is_open s) (hn : 1 ≤ n) : + continuous_on (λ x, fderiv 𝕜 f x) s := +((cont_diff_on_succ_iff_fderiv_of_open hs).1 (h.of_le hn)).2.continuous_on + +/-! ### Functions with a Taylor series on the whole space -/ + +/-- `has_ftaylor_series_up_to n f p` registers the fact that `p 0 = f` and `p (m+1)` is a +derivative of `p m` for `m < n`, and is continuous for `m ≤ n`. This is a predicate analogous to +`has_fderiv_at` but for higher order derivatives. -/ +structure has_ftaylor_series_up_to (n : ℕ∞) + (f : E → F) (p : E → formal_multilinear_series 𝕜 E F) : Prop := +(zero_eq : ∀ x, (p x 0).uncurry0 = f x) +(fderiv : ∀ (m : ℕ) (hm : (m : ℕ∞) < n), ∀ x, + has_fderiv_at (λ y, p y m) (p x m.succ).curry_left x) +(cont : ∀ (m : ℕ) (hm : (m : ℕ∞) ≤ n), continuous (λ x, p x m)) + +lemma has_ftaylor_series_up_to.zero_eq' + (h : has_ftaylor_series_up_to n f p) (x : E) : + p x 0 = (continuous_multilinear_curry_fin0 𝕜 E F).symm (f x) := +by { rw ← h.zero_eq x, symmetry, exact continuous_multilinear_map.uncurry0_curry0 _ } + +lemma has_ftaylor_series_up_to_on_univ_iff : + has_ftaylor_series_up_to_on n f p univ ↔ has_ftaylor_series_up_to n f p := +begin + split, + { assume H, + split, + { exact λ x, H.zero_eq x (mem_univ x) }, + { assume m hm x, + rw ← has_fderiv_within_at_univ, + exact H.fderiv_within m hm x (mem_univ x) }, + { assume m hm, + rw continuous_iff_continuous_on_univ, + exact H.cont m hm } }, + { assume H, + split, + { exact λ x hx, H.zero_eq x }, + { assume m hm x hx, + rw has_fderiv_within_at_univ, + exact H.fderiv m hm x }, + { assume m hm, + rw ← continuous_iff_continuous_on_univ, + exact H.cont m hm } } +end + +lemma has_ftaylor_series_up_to.has_ftaylor_series_up_to_on + (h : has_ftaylor_series_up_to n f p) (s : set E) : + has_ftaylor_series_up_to_on n f p s := +(has_ftaylor_series_up_to_on_univ_iff.2 h).mono (subset_univ _) + +lemma has_ftaylor_series_up_to.of_le + (h : has_ftaylor_series_up_to n f p) (hmn : m ≤ n) : + has_ftaylor_series_up_to m f p := +by { rw ← has_ftaylor_series_up_to_on_univ_iff at h ⊢, exact h.of_le hmn } + +lemma has_ftaylor_series_up_to.continuous + (h : has_ftaylor_series_up_to n f p) : continuous f := +begin + rw ← has_ftaylor_series_up_to_on_univ_iff at h, + rw continuous_iff_continuous_on_univ, + exact h.continuous_on +end + +lemma has_ftaylor_series_up_to_zero_iff : + has_ftaylor_series_up_to 0 f p ↔ continuous f ∧ (∀ x, (p x 0).uncurry0 = f x) := +by simp [has_ftaylor_series_up_to_on_univ_iff.symm, continuous_iff_continuous_on_univ, + has_ftaylor_series_up_to_on_zero_iff] + +/-- If a function has a Taylor series at order at least `1`, then the term of order `1` of this +series is a derivative of `f`. -/ +lemma has_ftaylor_series_up_to.has_fderiv_at + (h : has_ftaylor_series_up_to n f p) (hn : 1 ≤ n) (x : E) : + has_fderiv_at f (continuous_multilinear_curry_fin1 𝕜 E F (p x 1)) x := +begin + rw [← has_fderiv_within_at_univ], + exact (has_ftaylor_series_up_to_on_univ_iff.2 h).has_fderiv_within_at hn (mem_univ _) +end + +lemma has_ftaylor_series_up_to.differentiable + (h : has_ftaylor_series_up_to n f p) (hn : 1 ≤ n) : differentiable 𝕜 f := +λ x, (h.has_fderiv_at hn x).differentiable_at + +/-- `p` is a Taylor series of `f` up to `n+1` if and only if `p.shift` is a Taylor series up to `n` +for `p 1`, which is a derivative of `f`. -/ +theorem has_ftaylor_series_up_to_succ_iff_right {n : ℕ} : + has_ftaylor_series_up_to ((n + 1) : ℕ) f p ↔ + (∀ x, (p x 0).uncurry0 = f x) + ∧ (∀ x, has_fderiv_at (λ y, p y 0) (p x 1).curry_left x) + ∧ has_ftaylor_series_up_to n + (λ x, continuous_multilinear_curry_fin1 𝕜 E F (p x 1)) (λ x, (p x).shift) := +by simp only [has_ftaylor_series_up_to_on_succ_iff_right, ← has_ftaylor_series_up_to_on_univ_iff, + mem_univ, forall_true_left, has_fderiv_within_at_univ] + +/-! ### Smooth functions at a point -/ + +variable (𝕜) + +/-- A function is continuously differentiable up to `n` at a point `x` if, for any integer `k ≤ n`, +there is a neighborhood of `x` where `f` admits derivatives up to order `n`, which are continuous. +-/ +def cont_diff_at (n : ℕ∞) (f : E → F) (x : E) : Prop := +cont_diff_within_at 𝕜 n f univ x + +variable {𝕜} + +theorem cont_diff_within_at_univ : + cont_diff_within_at 𝕜 n f univ x ↔ cont_diff_at 𝕜 n f x := +iff.rfl + +lemma cont_diff_at_top : + cont_diff_at 𝕜 ∞ f x ↔ ∀ (n : ℕ), cont_diff_at 𝕜 n f x := +by simp [← cont_diff_within_at_univ, cont_diff_within_at_top] + +lemma cont_diff_at.cont_diff_within_at + (h : cont_diff_at 𝕜 n f x) : cont_diff_within_at 𝕜 n f s x := +h.mono (subset_univ _) + +lemma cont_diff_within_at.cont_diff_at + (h : cont_diff_within_at 𝕜 n f s x) (hx : s ∈ 𝓝 x) : + cont_diff_at 𝕜 n f x := +by rwa [cont_diff_at, ← cont_diff_within_at_inter hx, univ_inter] + +lemma cont_diff_at.congr_of_eventually_eq + (h : cont_diff_at 𝕜 n f x) (hg : f₁ =ᶠ[𝓝 x] f) : + cont_diff_at 𝕜 n f₁ x := +h.congr_of_eventually_eq' (by rwa nhds_within_univ) (mem_univ x) + +lemma cont_diff_at.of_le + (h : cont_diff_at 𝕜 n f x) (hmn : m ≤ n) : + cont_diff_at 𝕜 m f x := +h.of_le hmn + +lemma cont_diff_at.continuous_at + (h : cont_diff_at 𝕜 n f x) : continuous_at f x := +by simpa [continuous_within_at_univ] using h.continuous_within_at + +/-- If a function is `C^n` with `n ≥ 1` at a point, then it is differentiable there. -/ +lemma cont_diff_at.differentiable_at + (h : cont_diff_at 𝕜 n f x) (hn : 1 ≤ n) : differentiable_at 𝕜 f x := +by simpa [hn, differentiable_within_at_univ] using h.differentiable_within_at + +/-- A function is `C^(n + 1)` at a point iff locally, it has a derivative which is `C^n`. -/ +theorem cont_diff_at_succ_iff_has_fderiv_at {n : ℕ} : + cont_diff_at 𝕜 ((n + 1) : ℕ) f x + ↔ (∃ f' : E → E →L[𝕜] F, (∃ u ∈ 𝓝 x, ∀ x ∈ u, has_fderiv_at f (f' x) x) + ∧ cont_diff_at 𝕜 n f' x) := +begin + rw [← cont_diff_within_at_univ, cont_diff_within_at_succ_iff_has_fderiv_within_at], + simp only [nhds_within_univ, exists_prop, mem_univ, insert_eq_of_mem], + split, + { rintros ⟨u, H, f', h_fderiv, h_cont_diff⟩, + rcases mem_nhds_iff.mp H with ⟨t, htu, ht, hxt⟩, + refine ⟨f', ⟨t, _⟩, h_cont_diff.cont_diff_at H⟩, + refine ⟨mem_nhds_iff.mpr ⟨t, subset.rfl, ht, hxt⟩, _⟩, + intros y hyt, + refine (h_fderiv y (htu hyt)).has_fderiv_at _, + exact mem_nhds_iff.mpr ⟨t, htu, ht, hyt⟩ }, + { rintros ⟨f', ⟨u, H, h_fderiv⟩, h_cont_diff⟩, + refine ⟨u, H, f', _, h_cont_diff.cont_diff_within_at⟩, + intros x hxu, + exact (h_fderiv x hxu).has_fderiv_within_at } +end + +protected theorem cont_diff_at.eventually {n : ℕ} (h : cont_diff_at 𝕜 n f x) : + ∀ᶠ y in 𝓝 x, cont_diff_at 𝕜 n f y := +by simpa [nhds_within_univ] using h.eventually + +/-! ### Smooth functions -/ + +variable (𝕜) + +/-- A function is continuously differentiable up to `n` if it admits derivatives up to +order `n`, which are continuous. Contrary to the case of definitions in domains (where derivatives +might not be unique) we do not need to localize the definition in space or time. +-/ +def cont_diff (n : ℕ∞) (f : E → F) : Prop := +∃ p : E → formal_multilinear_series 𝕜 E F, has_ftaylor_series_up_to n f p + +variable {𝕜} + +theorem cont_diff_on_univ : cont_diff_on 𝕜 n f univ ↔ cont_diff 𝕜 n f := +begin + split, + { assume H, + use ftaylor_series_within 𝕜 f univ, + rw ← has_ftaylor_series_up_to_on_univ_iff, + exact H.ftaylor_series_within unique_diff_on_univ }, + { rintros ⟨p, hp⟩ x hx m hm, + exact ⟨univ, filter.univ_sets _, p, (hp.has_ftaylor_series_up_to_on univ).of_le hm⟩ } +end + +lemma cont_diff_iff_cont_diff_at : cont_diff 𝕜 n f ↔ ∀ x, cont_diff_at 𝕜 n f x := +by simp [← cont_diff_on_univ, cont_diff_on, cont_diff_at] + +lemma cont_diff.cont_diff_at (h : cont_diff 𝕜 n f) : cont_diff_at 𝕜 n f x := +cont_diff_iff_cont_diff_at.1 h x + +lemma cont_diff.cont_diff_within_at (h : cont_diff 𝕜 n f) : cont_diff_within_at 𝕜 n f s x := +h.cont_diff_at.cont_diff_within_at + +lemma cont_diff_top : cont_diff 𝕜 ∞ f ↔ ∀ (n : ℕ), cont_diff 𝕜 n f := +by simp [cont_diff_on_univ.symm, cont_diff_on_top] + +lemma cont_diff_all_iff_nat : (∀ n, cont_diff 𝕜 n f) ↔ (∀ n : ℕ, cont_diff 𝕜 n f) := +by simp only [← cont_diff_on_univ, cont_diff_on_all_iff_nat] + +lemma cont_diff.cont_diff_on (h : cont_diff 𝕜 n f) : cont_diff_on 𝕜 n f s := +(cont_diff_on_univ.2 h).mono (subset_univ _) + +@[simp] lemma cont_diff_zero : cont_diff 𝕜 0 f ↔ continuous f := +begin + rw [← cont_diff_on_univ, continuous_iff_continuous_on_univ], + exact cont_diff_on_zero +end + +lemma cont_diff_at_zero : cont_diff_at 𝕜 0 f x ↔ ∃ u ∈ 𝓝 x, continuous_on f u := +by { rw ← cont_diff_within_at_univ, simp [cont_diff_within_at_zero, nhds_within_univ] } + +theorem cont_diff_at_one_iff : cont_diff_at 𝕜 1 f x ↔ + ∃ f' : E → (E →L[𝕜] F), ∃ u ∈ 𝓝 x, continuous_on f' u ∧ ∀ x ∈ u, has_fderiv_at f (f' x) x := +by simp_rw [show (1 : ℕ∞) = (0 + 1 : ℕ), from (zero_add 1).symm, + cont_diff_at_succ_iff_has_fderiv_at, show ((0 : ℕ) : ℕ∞) = 0, from rfl, + cont_diff_at_zero, exists_mem_and_iff antitone_bforall antitone_continuous_on, and_comm] + +lemma cont_diff.of_le (h : cont_diff 𝕜 n f) (hmn : m ≤ n) : cont_diff 𝕜 m f := +cont_diff_on_univ.1 $ (cont_diff_on_univ.2 h).of_le hmn + +lemma cont_diff.of_succ {n : ℕ} (h : cont_diff 𝕜 (n + 1) f) : cont_diff 𝕜 n f := +h.of_le $ with_top.coe_le_coe.mpr le_self_add + +lemma cont_diff.one_of_succ {n : ℕ} (h : cont_diff 𝕜 (n + 1) f) : cont_diff 𝕜 1 f := +h.of_le $ with_top.coe_le_coe.mpr le_add_self + +lemma cont_diff.continuous (h : cont_diff 𝕜 n f) : continuous f := +cont_diff_zero.1 (h.of_le bot_le) + +/-- If a function is `C^n` with `n ≥ 1`, then it is differentiable. -/ +lemma cont_diff.differentiable (h : cont_diff 𝕜 n f) (hn : 1 ≤ n) : differentiable 𝕜 f := +differentiable_on_univ.1 $ (cont_diff_on_univ.2 h).differentiable_on hn + +lemma cont_diff_iff_forall_nat_le : + cont_diff 𝕜 n f ↔ ∀ m : ℕ, ↑m ≤ n → cont_diff 𝕜 m f := +by { simp_rw [← cont_diff_on_univ], exact cont_diff_on_iff_forall_nat_le } + + +/-! ### Iterated derivative -/ + +variable (𝕜) + +/-- The `n`-th derivative of a function, as a multilinear map, defined inductively. -/ +noncomputable def iterated_fderiv (n : ℕ) (f : E → F) : + E → (E [×n]→L[𝕜] F) := +nat.rec_on n + (λ x, continuous_multilinear_map.curry0 𝕜 E (f x)) + (λ n rec x, continuous_linear_map.uncurry_left (fderiv 𝕜 rec x)) + +/-- Formal Taylor series associated to a function within a set. -/ +def ftaylor_series (f : E → F) (x : E) : formal_multilinear_series 𝕜 E F := +λ n, iterated_fderiv 𝕜 n f x + +variable {𝕜} + +@[simp] lemma iterated_fderiv_zero_apply (m : (fin 0) → E) : + (iterated_fderiv 𝕜 0 f x : ((fin 0) → E) → F) m = f x := rfl + +lemma iterated_fderiv_zero_eq_comp : + iterated_fderiv 𝕜 0 f = (continuous_multilinear_curry_fin0 𝕜 E F).symm ∘ f := rfl + +@[simp] lemma norm_iterated_fderiv_zero : + ‖iterated_fderiv 𝕜 0 f x‖ = ‖f x‖ := +by rw [iterated_fderiv_zero_eq_comp, linear_isometry_equiv.norm_map] + +lemma iterated_fderiv_with_zero_eq : + iterated_fderiv_within 𝕜 0 f s = iterated_fderiv 𝕜 0 f := +by { ext, refl } + +lemma iterated_fderiv_succ_apply_left {n : ℕ} (m : fin (n + 1) → E): + (iterated_fderiv 𝕜 (n + 1) f x : (fin (n + 1) → E) → F) m + = (fderiv 𝕜 (iterated_fderiv 𝕜 n f) x : E → (E [×n]→L[𝕜] F)) (m 0) (tail m) := rfl + +/-- Writing explicitly the `n+1`-th derivative as the composition of a currying linear equiv, +and the derivative of the `n`-th derivative. -/ +lemma iterated_fderiv_succ_eq_comp_left {n : ℕ} : + iterated_fderiv 𝕜 (n + 1) f = + (continuous_multilinear_curry_left_equiv 𝕜 (λ (i : fin (n + 1)), E) F) + ∘ (fderiv 𝕜 (iterated_fderiv 𝕜 n f)) := rfl + +/-- Writing explicitly the derivative of the `n`-th derivative as the composition of a currying +linear equiv, and the `n + 1`-th derivative. -/ +lemma fderiv_iterated_fderiv {n : ℕ} : + fderiv 𝕜 (iterated_fderiv 𝕜 n f) = + (continuous_multilinear_curry_left_equiv 𝕜 (λ (i : fin (n + 1)), E) F).symm + ∘ (iterated_fderiv 𝕜 (n + 1) f) := +begin + rw iterated_fderiv_succ_eq_comp_left, + ext1 x, + simp only [function.comp_app, linear_isometry_equiv.symm_apply_apply], +end + +lemma has_compact_support.iterated_fderiv (hf : has_compact_support f) (n : ℕ) : + has_compact_support (iterated_fderiv 𝕜 n f) := +begin + induction n with n IH, + { rw [iterated_fderiv_zero_eq_comp], + apply hf.comp_left, + exact linear_isometry_equiv.map_zero _ }, + { rw iterated_fderiv_succ_eq_comp_left, + apply (IH.fderiv 𝕜).comp_left, + exact linear_isometry_equiv.map_zero _ } +end +lemma norm_fderiv_iterated_fderiv {n : ℕ} : + ‖fderiv 𝕜 (iterated_fderiv 𝕜 n f) x‖ = ‖iterated_fderiv 𝕜 (n + 1) f x‖ := +by rw [iterated_fderiv_succ_eq_comp_left, linear_isometry_equiv.norm_map] + +lemma iterated_fderiv_within_univ {n : ℕ} : + iterated_fderiv_within 𝕜 n f univ = iterated_fderiv 𝕜 n f := +begin + induction n with n IH, + { ext x, simp }, + { ext x m, + rw [iterated_fderiv_succ_apply_left, iterated_fderiv_within_succ_apply_left, IH, + fderiv_within_univ] } +end + +/-- In an open set, the iterated derivative within this set coincides with the global iterated +derivative. -/ +lemma iterated_fderiv_within_of_is_open (n : ℕ) (hs : is_open s) : + eq_on (iterated_fderiv_within 𝕜 n f s) (iterated_fderiv 𝕜 n f) s := +begin + induction n with n IH, + { assume x hx, + ext1 m, + simp only [iterated_fderiv_within_zero_apply, iterated_fderiv_zero_apply] }, + { assume x hx, + rw [iterated_fderiv_succ_eq_comp_left, iterated_fderiv_within_succ_eq_comp_left], + dsimp, + congr' 1, + rw fderiv_within_of_open hs hx, + apply filter.eventually_eq.fderiv_eq, + filter_upwards [hs.mem_nhds hx], + exact IH } +end + +lemma ftaylor_series_within_univ : + ftaylor_series_within 𝕜 f univ = ftaylor_series 𝕜 f := +begin + ext1 x, ext1 n, + change iterated_fderiv_within 𝕜 n f univ x = iterated_fderiv 𝕜 n f x, + rw iterated_fderiv_within_univ +end + +theorem iterated_fderiv_succ_apply_right {n : ℕ} (m : fin (n + 1) → E) : + (iterated_fderiv 𝕜 (n + 1) f x : (fin (n + 1) → E) → F) m + = iterated_fderiv 𝕜 n (λy, fderiv 𝕜 f y) x (init m) (m (last n)) := +begin + rw [← iterated_fderiv_within_univ, ← iterated_fderiv_within_univ, ← fderiv_within_univ], + exact iterated_fderiv_within_succ_apply_right unique_diff_on_univ (mem_univ _) _ +end + +/-- Writing explicitly the `n+1`-th derivative as the composition of a currying linear equiv, +and the `n`-th derivative of the derivative. -/ +lemma iterated_fderiv_succ_eq_comp_right {n : ℕ} : + iterated_fderiv 𝕜 (n + 1) f x = + ((continuous_multilinear_curry_right_equiv' 𝕜 n E F) + ∘ (iterated_fderiv 𝕜 n (λy, fderiv 𝕜 f y))) x := +by { ext m, rw iterated_fderiv_succ_apply_right, refl } + +lemma norm_iterated_fderiv_fderiv {n : ℕ} : + ‖iterated_fderiv 𝕜 n (fderiv 𝕜 f) x‖ = ‖iterated_fderiv 𝕜 (n + 1) f x‖ := +by rw [iterated_fderiv_succ_eq_comp_right, linear_isometry_equiv.norm_map] + +@[simp] lemma iterated_fderiv_one_apply (m : (fin 1) → E) : + (iterated_fderiv 𝕜 1 f x : ((fin 1) → E) → F) m + = (fderiv 𝕜 f x : E → F) (m 0) := +by { rw [iterated_fderiv_succ_apply_right, iterated_fderiv_zero_apply], refl } + +/-- When a function is `C^n` in a set `s` of unique differentiability, it admits +`ftaylor_series_within 𝕜 f s` as a Taylor series up to order `n` in `s`. -/ +theorem cont_diff_iff_ftaylor_series : + cont_diff 𝕜 n f ↔ has_ftaylor_series_up_to n f (ftaylor_series 𝕜 f) := +begin + split, + { rw [← cont_diff_on_univ, ← has_ftaylor_series_up_to_on_univ_iff, + ← ftaylor_series_within_univ], + exact λ h, cont_diff_on.ftaylor_series_within h unique_diff_on_univ }, + { assume h, exact ⟨ftaylor_series 𝕜 f, h⟩ } +end + +lemma cont_diff_iff_continuous_differentiable : + cont_diff 𝕜 n f ↔ + (∀ (m : ℕ), (m : ℕ∞) ≤ n → continuous (λ x, iterated_fderiv 𝕜 m f x)) + ∧ (∀ (m : ℕ), (m : ℕ∞) < n → differentiable 𝕜 (λ x, iterated_fderiv 𝕜 m f x)) := +by simp [cont_diff_on_univ.symm, continuous_iff_continuous_on_univ, + differentiable_on_univ.symm, iterated_fderiv_within_univ, + cont_diff_on_iff_continuous_on_differentiable_on unique_diff_on_univ] + +/-- If `f` is `C^n` then its `m`-times iterated derivative is continuous for `m ≤ n`. -/ +lemma cont_diff.continuous_iterated_fderiv {m : ℕ} (hm : (m : ℕ∞) ≤ n) + (hf : cont_diff 𝕜 n f) : continuous (λ x, iterated_fderiv 𝕜 m f x) := +(cont_diff_iff_continuous_differentiable.mp hf).1 m hm + +/-- If `f` is `C^n` then its `m`-times iterated derivative is differentiable for `m < n`. -/ +lemma cont_diff.differentiable_iterated_fderiv {m : ℕ} (hm : (m : ℕ∞) < n) + (hf : cont_diff 𝕜 n f) : differentiable 𝕜 (λ x, iterated_fderiv 𝕜 m f x) := +(cont_diff_iff_continuous_differentiable.mp hf).2 m hm + +lemma cont_diff_of_differentiable_iterated_fderiv + (h : ∀(m : ℕ), (m : ℕ∞) ≤ n → differentiable 𝕜 (iterated_fderiv 𝕜 m f)) : + cont_diff 𝕜 n f := +cont_diff_iff_continuous_differentiable.2 +⟨λ m hm, (h m hm).continuous, λ m hm, (h m (le_of_lt hm))⟩ + +/-- A function is `C^(n + 1)` if and only if it is differentiable, +and its derivative (formulated in terms of `fderiv`) is `C^n`. -/ +theorem cont_diff_succ_iff_fderiv {n : ℕ} : + cont_diff 𝕜 ((n + 1) : ℕ) f ↔ + differentiable 𝕜 f ∧ cont_diff 𝕜 n (λ y, fderiv 𝕜 f y) := +by simp only [← cont_diff_on_univ, ← differentiable_on_univ, ← fderiv_within_univ, + cont_diff_on_succ_iff_fderiv_within unique_diff_on_univ] + +theorem cont_diff_one_iff_fderiv : + cont_diff 𝕜 1 f ↔ differentiable 𝕜 f ∧ continuous (fderiv 𝕜 f) := +cont_diff_succ_iff_fderiv.trans $ iff.rfl.and cont_diff_zero + +/-- A function is `C^∞` if and only if it is differentiable, +and its derivative (formulated in terms of `fderiv`) is `C^∞`. -/ +theorem cont_diff_top_iff_fderiv : + cont_diff 𝕜 ∞ f ↔ + differentiable 𝕜 f ∧ cont_diff 𝕜 ∞ (λ y, fderiv 𝕜 f y) := +begin + simp only [← cont_diff_on_univ, ← differentiable_on_univ, ← fderiv_within_univ], + rw cont_diff_on_top_iff_fderiv_within unique_diff_on_univ, +end + +lemma cont_diff.continuous_fderiv + (h : cont_diff 𝕜 n f) (hn : 1 ≤ n) : + continuous (λ x, fderiv 𝕜 f x) := +((cont_diff_succ_iff_fderiv).1 (h.of_le hn)).2.continuous + +/-- If a function is at least `C^1`, its bundled derivative (mapping `(x, v)` to `Df(x) v`) is +continuous. -/ +lemma cont_diff.continuous_fderiv_apply + (h : cont_diff 𝕜 n f) (hn : 1 ≤ n) : + continuous (λp : E × E, (fderiv 𝕜 f p.1 : E → F) p.2) := +have A : continuous (λq : (E →L[𝕜] F) × E, q.1 q.2) := is_bounded_bilinear_map_apply.continuous, +have B : continuous (λp : E × E, (fderiv 𝕜 f p.1, p.2)) := + ((h.continuous_fderiv hn).comp continuous_fst).prod_mk continuous_snd, +A.comp B From c085f3044fe585c575e322bfab45b3633c48d820 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Wed, 22 Mar 2023 06:13:18 +0000 Subject: [PATCH 0670/1291] chore(*): add mathlib4 synchronization comments (#18630) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.category.Mon.basic` * `algebra.squarefree` * `category_theory.closed.monoidal` * `category_theory.single_obj` * `data.zmod.coprime` * `field_theory.perfect_closure` * `linear_algebra.invariant_basis_number` * `ring_theory.euclidean_domain` * `ring_theory.flat` * `ring_theory.int.basic` * `ring_theory.principal_ideal_domain` * `ring_theory.unique_factorization_domain` --- src/algebra/category/Mon/basic.lean | 3 +++ src/algebra/squarefree.lean | 3 +++ src/category_theory/closed/monoidal.lean | 3 +++ src/category_theory/single_obj.lean | 3 +++ src/data/zmod/coprime.lean | 3 +++ src/field_theory/perfect_closure.lean | 3 +++ src/linear_algebra/invariant_basis_number.lean | 3 +++ src/ring_theory/euclidean_domain.lean | 3 +++ src/ring_theory/flat.lean | 3 +++ src/ring_theory/int/basic.lean | 3 +++ src/ring_theory/principal_ideal_domain.lean | 3 +++ src/ring_theory/unique_factorization_domain.lean | 3 +++ 12 files changed, 36 insertions(+) diff --git a/src/algebra/category/Mon/basic.lean b/src/algebra/category/Mon/basic.lean index e83789a6b8d17..5b050d6a7f05f 100644 --- a/src/algebra/category/Mon/basic.lean +++ b/src/algebra/category/Mon/basic.lean @@ -10,6 +10,9 @@ import category_theory.functor.reflects_isomorphisms /-! # Category instances for monoid, add_monoid, comm_monoid, and add_comm_monoid. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We introduce the bundled categories: * `Mon` * `AddMon` diff --git a/src/algebra/squarefree.lean b/src/algebra/squarefree.lean index 8f7611aa8d9a4..fe3338803a6b6 100644 --- a/src/algebra/squarefree.lean +++ b/src/algebra/squarefree.lean @@ -7,6 +7,9 @@ import ring_theory.unique_factorization_domain /-! # Squarefree elements of monoids + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. An element of a monoid is squarefree when it is not divisible by any squares except the squares of units. diff --git a/src/category_theory/closed/monoidal.lean b/src/category_theory/closed/monoidal.lean index f583bab91441f..57f1979923cfb 100644 --- a/src/category_theory/closed/monoidal.lean +++ b/src/category_theory/closed/monoidal.lean @@ -11,6 +11,9 @@ import category_theory.functor.inv_isos /-! # Closed monoidal categories +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Define (right) closed objects and (right) closed monoidal categories. ## TODO diff --git a/src/category_theory/single_obj.lean b/src/category_theory/single_obj.lean index 082a39641d817..43eaf45676791 100644 --- a/src/category_theory/single_obj.lean +++ b/src/category_theory/single_obj.lean @@ -11,6 +11,9 @@ import combinatorics.quiver.single_obj /-! # Single-object category +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Single object category with a given monoid of endomorphisms. It is defined to facilitate transfering some definitions and lemmas (e.g., conjugacy etc.) from category theory to monoids and groups. diff --git a/src/data/zmod/coprime.lean b/src/data/zmod/coprime.lean index e1bee15b7c3da..967b2ffe4df5b 100644 --- a/src/data/zmod/coprime.lean +++ b/src/data/zmod/coprime.lean @@ -9,6 +9,9 @@ import ring_theory.int.basic /-! # Coprimality and vanishing +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We show that for prime `p`, the image of an integer `a` in `zmod p` vanishes if and only if `a` and `p` are not coprime. -/ diff --git a/src/field_theory/perfect_closure.lean b/src/field_theory/perfect_closure.lean index 78b1dae5c0ca0..fd21739bf06bc 100644 --- a/src/field_theory/perfect_closure.lean +++ b/src/field_theory/perfect_closure.lean @@ -9,6 +9,9 @@ import algebra.ring.equiv /-! # The perfect closure of a field + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ universes u v diff --git a/src/linear_algebra/invariant_basis_number.lean b/src/linear_algebra/invariant_basis_number.lean index 27204aca6438a..ef9b54e5dcfbe 100644 --- a/src/linear_algebra/invariant_basis_number.lean +++ b/src/linear_algebra/invariant_basis_number.lean @@ -9,6 +9,9 @@ import ring_theory.principal_ideal_domain /-! # Invariant basis number property +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We say that a ring `R` satisfies the invariant basis number property if there is a well-defined notion of the rank of a finitely generated free (left) `R`-module. Since a finitely generated free module with a basis consisting of `n` elements is linearly equivalent to `fin n → R`, it is diff --git a/src/ring_theory/euclidean_domain.lean b/src/ring_theory/euclidean_domain.lean index 9be725299613f..7ed33ee3b4815 100644 --- a/src/ring_theory/euclidean_domain.lean +++ b/src/ring_theory/euclidean_domain.lean @@ -11,6 +11,9 @@ import ring_theory.principal_ideal_domain /-! # Lemmas about Euclidean domains +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Various about Euclidean domains are proved; all of them seem to be true more generally for principal ideal domains, so these lemmas should probably be reproved in more generality and this file perhaps removed? diff --git a/src/ring_theory/flat.lean b/src/ring_theory/flat.lean index 802c2fac29e21..ffce2a604777c 100644 --- a/src/ring_theory/flat.lean +++ b/src/ring_theory/flat.lean @@ -9,6 +9,9 @@ import ring_theory.noetherian /-! # Flat modules +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A module `M` over a commutative ring `R` is *flat* if for all finitely generated ideals `I` of `R`, the canonical map `I ⊗ M →ₗ M` is injective. diff --git a/src/ring_theory/int/basic.lean b/src/ring_theory/int/basic.lean index 76f1f774b03a0..086a97a34cc12 100644 --- a/src/ring_theory/int/basic.lean +++ b/src/ring_theory/int/basic.lean @@ -11,6 +11,9 @@ import ring_theory.principal_ideal_domain /-! # Divisibility over ℕ and ℤ +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file collects results for the integers and natural numbers that use abstract algebra in their proofs or cases of ℕ and ℤ being examples of structures in abstract algebra. diff --git a/src/ring_theory/principal_ideal_domain.lean b/src/ring_theory/principal_ideal_domain.lean index fb6fdcf722a6c..4b08be8b78890 100644 --- a/src/ring_theory/principal_ideal_domain.lean +++ b/src/ring_theory/principal_ideal_domain.lean @@ -9,6 +9,9 @@ import ring_theory.unique_factorization_domain /-! # Principal ideal rings and principal ideal domains +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A principal ideal ring (PIR) is a ring in which all left ideals are principal. A principal ideal domain (PID) is an integral domain which is a principal ideal ring. diff --git a/src/ring_theory/unique_factorization_domain.lean b/src/ring_theory/unique_factorization_domain.lean index 96f1548a897b8..471af7ef520d2 100644 --- a/src/ring_theory/unique_factorization_domain.lean +++ b/src/ring_theory/unique_factorization_domain.lean @@ -14,6 +14,9 @@ import ring_theory.multiplicity # Unique factorization +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main Definitions * `wf_dvd_monoid` holds for `monoid`s for which a strict divisibility relation is well-founded. From 57e09a1296bfb4330ddf6624f1028ba186117d82 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 22 Mar 2023 10:05:48 +0000 Subject: [PATCH 0671/1291] feat(algebra/monoid_algebra): add division by a generator (#15905) This generalizes `polynomial.div_X`. [Zulip thread](https://leanprover.zulipchat.com/#narrow/stream/217875-Is-there-code-for-X.3F/topic/X_i.20divides.20f.20if.20f.3D0.20when.20X_i.3D0/near/339570367) --- src/algebra/monoid_algebra/basic.lean | 40 +++++ src/algebra/monoid_algebra/division.lean | 180 +++++++++++++++++++++++ src/data/mv_polynomial/division.lean | 154 +++++++++++++++++++ src/data/polynomial/inductions.lean | 12 +- 4 files changed, 378 insertions(+), 8 deletions(-) create mode 100644 src/algebra/monoid_algebra/division.lean create mode 100644 src/data/mv_polynomial/division.lean diff --git a/src/algebra/monoid_algebra/basic.lean b/src/algebra/monoid_algebra/basic.lean index d68da00ac395c..f746821df66ff 100644 --- a/src/algebra/monoid_algebra/basic.lean +++ b/src/algebra/monoid_algebra/basic.lean @@ -433,6 +433,21 @@ lemma mul_single_one_apply [mul_one_class G] (f : monoid_algebra k G) (r : k) (x (f * single 1 r) x = f x * r := f.mul_single_apply_aux $ λ a, by rw [mul_one] +lemma mul_single_apply_of_not_exists_mul [has_mul G] (r : k) {g g' : G} (x : monoid_algebra k G) + (h : ¬∃ d, g' = d * g): + (x * finsupp.single g r : monoid_algebra k G) g' = 0 := +begin + classical, + rw [mul_apply, finsupp.sum_comm, finsupp.sum_single_index], + swap, + { simp_rw [finsupp.sum, mul_zero, if_t_t, finset.sum_const_zero] }, + { apply finset.sum_eq_zero, + simp_rw ite_eq_right_iff, + rintros g'' hg'' rfl, + exfalso, + exact h ⟨_, rfl⟩ } +end + lemma single_mul_apply_aux [has_mul G] (f : monoid_algebra k G) {r : k} {x y z : G} (H : ∀ a, x * a = y ↔ a = z) : (single x r * f) y = r * f z := @@ -448,6 +463,21 @@ lemma single_one_mul_apply [mul_one_class G] (f : monoid_algebra k G) (r : k) (x (single 1 r * f) x = r * f x := f.single_mul_apply_aux $ λ a, by rw [one_mul] +lemma single_mul_apply_of_not_exists_mul [has_mul G] (r : k) {g g' : G} (x : monoid_algebra k G) + (h : ¬∃ d, g' = g * d): + (finsupp.single g r * x : monoid_algebra k G) g' = 0 := +begin + classical, + rw [mul_apply, finsupp.sum_single_index], + swap, + { simp_rw [finsupp.sum, zero_mul, if_t_t, finset.sum_const_zero] }, + { apply finset.sum_eq_zero, + simp_rw ite_eq_right_iff, + rintros g'' hg'' rfl, + exfalso, + exact h ⟨_, rfl⟩ }, +end + lemma lift_nc_smul [mul_one_class G] {R : Type*} [semiring R] (f : k →+* R) (g : G →* R) (c : k) (φ : monoid_algebra k G) : lift_nc (f : k →+ R) g (c • φ) = f c * lift_nc (f : k →+ R) g φ := @@ -1224,6 +1254,11 @@ lemma mul_single_zero_apply [add_zero_class G] (f : add_monoid_algebra k G) (r : (f * single 0 r) x = f x * r := f.mul_single_apply_aux r _ _ _ $ λ a, by rw [add_zero] +lemma mul_single_apply_of_not_exists_add [has_add G] (r : k) {g g' : G} (x : add_monoid_algebra k G) + (h : ¬∃ d, g' = d + g): + (x * finsupp.single g r : add_monoid_algebra k G) g' = 0 := +@monoid_algebra.mul_single_apply_of_not_exists_mul k (multiplicative G) _ _ _ _ _ _ h + lemma single_mul_apply_aux [has_add G] (f : add_monoid_algebra k G) (r : k) (x y z : G) (H : ∀ a, x + a = y ↔ a = z) : (single x r * f : add_monoid_algebra k G) y = r * f z := @@ -1233,6 +1268,11 @@ lemma single_zero_mul_apply [add_zero_class G] (f : add_monoid_algebra k G) (r : (single 0 r * f : add_monoid_algebra k G) x = r * f x := f.single_mul_apply_aux r _ _ _ $ λ a, by rw [zero_add] +lemma single_mul_apply_of_not_exists_add [has_add G] (r : k) {g g' : G} (x : add_monoid_algebra k G) + (h : ¬∃ d, g' = g + d): + (finsupp.single g r * x : add_monoid_algebra k G) g' = 0 := +@monoid_algebra.single_mul_apply_of_not_exists_mul k (multiplicative G) _ _ _ _ _ _ h + lemma mul_single_apply [add_group G] (f : add_monoid_algebra k G) (r : k) (x y : G) : (f * single x r) y = f (y - x) * r := (sub_eq_add_neg y x).symm ▸ diff --git a/src/algebra/monoid_algebra/division.lean b/src/algebra/monoid_algebra/division.lean new file mode 100644 index 0000000000000..dcbd98f94a9f8 --- /dev/null +++ b/src/algebra/monoid_algebra/division.lean @@ -0,0 +1,180 @@ +/- +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.monoid_algebra.basic +import data.finsupp.order + +/-! +# Division of `add_monoid_algebra` by monomials + +This file is most important for when `G = ℕ` (polynomials) or `G = σ →₀ ℕ` (multivariate +polynomials). + +In order to apply in maximal generality (such as for `laurent_polynomial`s), this uses +`∃ d, g' = g + d` in many places instead of `g ≤ g'`. + +## Main definitions + +* `add_monoid_algebra.div_of x g`: divides `x` by the monomial `add_monoid_algebra.of k G g` +* `add_monoid_algebra.mod_of x g`: the remainder upon dividing `x` by the monomial + `add_monoid_algebra.of k G g`. + +## Main results + +* `add_monoid_algebra.div_of_add_mod_of`, `add_monoid_algebra.mod_of_add_div_of`: `div_of` and + `mod_of` are well-behaved as quotient and remainder operators. + +## Implementation notes + +`∃ d, g' = g + d` is used as opposed to some other permutation up to commutativity in order to match +the definition of `semigroup_has_dvd`. The results in this file could be duplicated for +`monoid_algebra` by using `g ∣ g'`, but this can't be done automatically, and in any case is not +likely to be very useful. + +-/ + + +variables {k G : Type*} [semiring k] + +namespace add_monoid_algebra + +section +variables [add_cancel_comm_monoid G] + +/-- Divide by `of' k G g`, discarding terms not divisible by this. -/ +noncomputable def div_of (x : add_monoid_algebra k G) (g : G) : add_monoid_algebra k G := +-- note: comapping by `+ g` has the effect of subtracting `g` from every element in the support, and +-- discarding the elements of the support from which `g` can't be subtracted. If `G` is an additive +-- group, such as `ℤ` when used for `laurent_polynomial`, then no discarding occurs. +@finsupp.comap_domain.add_monoid_hom _ _ _ _ ((+) g) + (add_right_injective g) x + +local infix ` /ᵒᶠ `:70 := div_of + +@[simp] lemma div_of_apply (g : G) (x : add_monoid_algebra k G) (g' : G) : + (x /ᵒᶠ g) g' = x (g + g') := rfl + +@[simp] lemma support_div_of (g : G) (x : add_monoid_algebra k G) : + (x /ᵒᶠ g).support = x.support.preimage ((+) g) + (function.injective.inj_on + (add_right_injective g) _) := rfl + +@[simp] lemma zero_div_of (g : G) : (0 : add_monoid_algebra k G) /ᵒᶠ g = 0 := +map_zero _ + +@[simp] lemma div_of_zero (x : add_monoid_algebra k G) : x /ᵒᶠ 0 = x := +by { ext, simp only [add_monoid_algebra.div_of_apply, zero_add] } + +lemma add_div_of (x y : add_monoid_algebra k G) (g : G) : (x + y) /ᵒᶠ g = x /ᵒᶠ g + y /ᵒᶠ g := +map_add _ _ _ + +lemma div_of_add (x : add_monoid_algebra k G) (a b : G) : + x /ᵒᶠ (a + b) = (x /ᵒᶠ a) /ᵒᶠ b := +by { ext, simp only [add_monoid_algebra.div_of_apply, add_assoc] } + +/-- A bundled version of `add_monoid_algebra.div_of`. -/ +@[simps] +noncomputable def div_of_hom : multiplicative G →* add_monoid.End (add_monoid_algebra k G) := +{ to_fun := λ g, + { to_fun := λ x, div_of x g.to_add, + map_zero' := zero_div_of _, + map_add' := λ x y, add_div_of x y g.to_add }, + map_one' := add_monoid_hom.ext div_of_zero, + map_mul' := λ g₁ g₂, add_monoid_hom.ext $ λ x, + (congr_arg _ (add_comm g₁.to_add g₂.to_add)).trans (div_of_add _ _ _) } + +lemma of'_mul_div_of (a : G) (x : add_monoid_algebra k G) : + (of' k G a * x) /ᵒᶠ a = x := +begin + ext b, + rw [add_monoid_algebra.div_of_apply, of'_apply, single_mul_apply_aux, one_mul], + intro c, + exact add_right_inj _, +end + +lemma mul_of'_div_of (x : add_monoid_algebra k G) (a : G) : + (x * of' k G a) /ᵒᶠ a = x := +begin + ext b, + rw [add_monoid_algebra.div_of_apply, of'_apply, mul_single_apply_aux, mul_one], + intro c, + rw add_comm, + exact add_right_inj _, +end + +lemma of'_div_of (a : G) : (of' k G a) /ᵒᶠ a = 1 := +by simpa only [one_mul] using mul_of'_div_of (1 : add_monoid_algebra k G) a + +/-- The remainder upon division by `of' k G g`. -/ +noncomputable def mod_of (x : add_monoid_algebra k G) (g : G) : add_monoid_algebra k G := +x.filter (λ g₁, ¬∃ g₂, g₁ = g + g₂) + +local infix ` %ᵒᶠ `:70 := mod_of + +@[simp] lemma mod_of_apply_of_not_exists_add (x : add_monoid_algebra k G) (g : G) (g' : G) + (h : ¬∃ d, g' = g + d) : + (x %ᵒᶠ g) g' = x g' := +finsupp.filter_apply_pos _ _ h + +@[simp] lemma mod_of_apply_of_exists_add (x : add_monoid_algebra k G) (g : G) (g' : G) + (h : ∃ d, g' = g + d) : + (x %ᵒᶠ g) g' = 0 := +finsupp.filter_apply_neg _ _ $ by rwa [not_not] + +@[simp] lemma mod_of_apply_add_self (x : add_monoid_algebra k G) (g : G) (d : G) : + (x %ᵒᶠ g) (d + g) = 0 := +mod_of_apply_of_exists_add _ _ _ ⟨_, add_comm _ _⟩ + +@[simp] lemma mod_of_apply_self_add (x : add_monoid_algebra k G) (g : G) (d : G) : + (x %ᵒᶠ g) (g + d) = 0 := +mod_of_apply_of_exists_add _ _ _ ⟨_, rfl⟩ + +lemma of'_mul_mod_of (g : G) (x : add_monoid_algebra k G) : + (of' k G g * x) %ᵒᶠ g = 0 := +begin + ext g', + rw finsupp.zero_apply, + obtain ⟨d, rfl⟩ | h := em (∃ d, g' = g + d), + { rw mod_of_apply_self_add }, + { rw [mod_of_apply_of_not_exists_add _ _ _ h, of'_apply, + single_mul_apply_of_not_exists_add _ _ h] }, +end + +lemma mul_of'_mod_of (x : add_monoid_algebra k G) (g : G) : + (x * of' k G g) %ᵒᶠ g = 0 := +begin + ext g', + rw finsupp.zero_apply, + obtain ⟨d, rfl⟩ | h := em (∃ d, g' = g + d), + { rw mod_of_apply_self_add }, + { rw [mod_of_apply_of_not_exists_add _ _ _ h, of'_apply, mul_single_apply_of_not_exists_add], + simpa only [add_comm] using h }, +end + +lemma of'_mod_of (g : G) : of' k G g %ᵒᶠ g = 0 := +by simpa only [one_mul] using mul_of'_mod_of (1 : add_monoid_algebra k G) g + +lemma div_of_add_mod_of (x : add_monoid_algebra k G) (g : G) : + of' k G g * (x /ᵒᶠ g) + x %ᵒᶠ g = x := +begin + ext g', + simp_rw [finsupp.add_apply], + obtain ⟨d, rfl⟩ | h := em (∃ d, g' = g + d), + swap, + { rw [mod_of_apply_of_not_exists_add _ _ _ h, of'_apply, single_mul_apply_of_not_exists_add _ _ h, + zero_add] }, + { rw [mod_of_apply_self_add, add_zero], + rw [of'_apply, single_mul_apply_aux _ _ _, one_mul, div_of_apply], + intro a, + exact add_right_inj _ } +end + +lemma mod_of_add_div_of (x : add_monoid_algebra k G) (g : G) : + x %ᵒᶠ g + of' k G g * (x /ᵒᶠ g) = x := +by rw [add_comm, div_of_add_mod_of] + +end + +end add_monoid_algebra diff --git a/src/data/mv_polynomial/division.lean b/src/data/mv_polynomial/division.lean new file mode 100644 index 0000000000000..3dfa2febc6cba --- /dev/null +++ b/src/data/mv_polynomial/division.lean @@ -0,0 +1,154 @@ +/- +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.monoid_algebra.division +import data.mv_polynomial.basic + +/-! +# Division of `mv_polynomial` by monomials + +## Main definitions + +* `mv_polynomial.div_monomial x s`: divides `x` by the monomial `mv_polynomial.monomial 1 s` +* `mv_polynomial.mod_monomial x s`: the remainder upon dividing `x` by the monomial + `mv_polynomial.monomial 1 s`. + +## Main results + +* `mv_polynomial.div_monomial_add_mod_monomial`, `mv_polynomial.mod_monomial_add_div_monomial`: + `div_monomial` and `mod_monomial` are well-behaved as quotient and remainder operators. + +## Implementation notes + +Where possible, the results in this file should be first proved in the generality of +`add_monoid_algebra`, and then the versions specialized to `mv_polynomial` proved in terms of these. + +-/ + + +variables {σ R : Type*} [comm_semiring R] + +namespace mv_polynomial + +section copied_declarations +/-! Please ensure the declarations in this section are direct translations of `add_monoid_algebra` +results. -/ + +/-- Divide by `monomial 1 s`, discarding terms not divisible by this. -/ +noncomputable def div_monomial (p : mv_polynomial σ R) (s : σ →₀ ℕ) : mv_polynomial σ R := +add_monoid_algebra.div_of p s + +local infix ` /ᵐᵒⁿᵒᵐⁱᵃˡ `:70 := div_monomial + +@[simp] lemma coeff_div_monomial (s : σ →₀ ℕ) (x : mv_polynomial σ R) (s' : σ →₀ ℕ) : + coeff s' (x /ᵐᵒⁿᵒᵐⁱᵃˡ s) = coeff (s + s') x := rfl + +@[simp] lemma support_div_monomial (s : σ →₀ ℕ) (x : mv_polynomial σ R) : + (x /ᵐᵒⁿᵒᵐⁱᵃˡ s).support = x.support.preimage _ ((add_right_injective s).inj_on _) := rfl + +@[simp] lemma zero_div_monomial (s : σ →₀ ℕ) : (0 : mv_polynomial σ R) /ᵐᵒⁿᵒᵐⁱᵃˡ s = 0 := +add_monoid_algebra.zero_div_of _ + +lemma div_monomial_zero (x : mv_polynomial σ R) : x /ᵐᵒⁿᵒᵐⁱᵃˡ 0 = x := +x.div_of_zero + +lemma add_div_monomial (x y : mv_polynomial σ R) (s : σ →₀ ℕ) : + (x + y) /ᵐᵒⁿᵒᵐⁱᵃˡ s = x /ᵐᵒⁿᵒᵐⁱᵃˡ s + y /ᵐᵒⁿᵒᵐⁱᵃˡ s := +map_add _ _ _ + +lemma div_monomial_add (a b : σ →₀ ℕ) (x : mv_polynomial σ R) : + x /ᵐᵒⁿᵒᵐⁱᵃˡ (a + b) = (x /ᵐᵒⁿᵒᵐⁱᵃˡ a) /ᵐᵒⁿᵒᵐⁱᵃˡ b := +x.div_of_add _ _ + +@[simp] lemma div_monomial_monomial_mul (a : σ →₀ ℕ) (x : mv_polynomial σ R) : + (monomial a 1 * x) /ᵐᵒⁿᵒᵐⁱᵃˡ a = x := +x.of'_mul_div_of _ + +@[simp] lemma div_monomial_mul_monomial (a : σ →₀ ℕ) (x : mv_polynomial σ R) : + (x * monomial a 1) /ᵐᵒⁿᵒᵐⁱᵃˡ a = x := +x.mul_of'_div_of _ + +@[simp] lemma div_monomial_monomial (a : σ →₀ ℕ) : + (monomial a 1) /ᵐᵒⁿᵒᵐⁱᵃˡ a = (1 : mv_polynomial σ R) := +add_monoid_algebra.of'_div_of _ + +/-- The remainder upon division by `monomial 1 s`. -/ +noncomputable def mod_monomial (x : mv_polynomial σ R) (s : σ →₀ ℕ) : mv_polynomial σ R := +x.mod_of s + +local infix ` %ᵐᵒⁿᵒᵐⁱᵃˡ `:70 := mod_monomial + +@[simp] lemma coeff_mod_monomial_of_not_le {s' s : σ →₀ ℕ} (x : mv_polynomial σ R) (h : ¬s ≤ s') : + coeff s' (x %ᵐᵒⁿᵒᵐⁱᵃˡ s) = coeff s' x := +x.mod_of_apply_of_not_exists_add s s' begin + rintro ⟨d, rfl⟩, + exact h le_self_add, +end + +@[simp] lemma coeff_mod_monomial_of_le {s' s : σ →₀ ℕ} (x : mv_polynomial σ R) (h : s ≤ s') : + coeff s' (x %ᵐᵒⁿᵒᵐⁱᵃˡ s) = 0 := +x.mod_of_apply_of_exists_add _ _ $ exists_add_of_le h + +@[simp] lemma monomial_mul_mod_monomial (s : σ →₀ ℕ) (x : mv_polynomial σ R) : + (monomial s 1 * x) %ᵐᵒⁿᵒᵐⁱᵃˡ s = 0 := +x.of'_mul_mod_of _ + +@[simp] lemma mul_monomial_mod_monomial (s : σ →₀ ℕ) (x : mv_polynomial σ R): + (x * monomial s 1) %ᵐᵒⁿᵒᵐⁱᵃˡ s = 0 := +x.mul_of'_mod_of _ + +@[simp] lemma monomial_mod_monomial (s : σ →₀ ℕ) : (monomial s (1 : R)) %ᵐᵒⁿᵒᵐⁱᵃˡ s = 0 := +add_monoid_algebra.of'_mod_of _ + +lemma div_monomial_add_mod_monomial (x : mv_polynomial σ R) (s : σ →₀ ℕ) : + monomial s 1 * (x /ᵐᵒⁿᵒᵐⁱᵃˡ s) + x %ᵐᵒⁿᵒᵐⁱᵃˡ s = x := +add_monoid_algebra.div_of_add_mod_of x s + +lemma mod_monomial_add_div_monomial (x : mv_polynomial σ R) (s : σ →₀ ℕ) : + x %ᵐᵒⁿᵒᵐⁱᵃˡ s + monomial s 1 * (x /ᵐᵒⁿᵒᵐⁱᵃˡ s) = x := +add_monoid_algebra.mod_of_add_div_of x s + +end copied_declarations + +section X_lemmas + +local infix ` /ᵐᵒⁿᵒᵐⁱᵃˡ `:70 := div_monomial +local infix ` %ᵐᵒⁿᵒᵐⁱᵃˡ `:70 := mod_monomial + +@[simp] lemma X_mul_div_monomial (i : σ) (x : mv_polynomial σ R) : + (X i * x) /ᵐᵒⁿᵒᵐⁱᵃˡ (finsupp.single i 1) = x := +div_monomial_monomial_mul _ _ + +@[simp] lemma X_div_monomial (i : σ) : + (X i : mv_polynomial σ R) /ᵐᵒⁿᵒᵐⁱᵃˡ (finsupp.single i 1) = 1 := +(div_monomial_monomial (finsupp.single i 1)) + +@[simp] lemma mul_X_div_monomial (x : mv_polynomial σ R) (i : σ) : + (x * X i) /ᵐᵒⁿᵒᵐⁱᵃˡ (finsupp.single i 1) = x := +div_monomial_mul_monomial _ _ + +@[simp] lemma X_mul_mod_monomial (i : σ) (x : mv_polynomial σ R) : + (X i * x) %ᵐᵒⁿᵒᵐⁱᵃˡ (finsupp.single i 1) = 0 := +monomial_mul_mod_monomial _ _ + +@[simp] lemma mul_X_mod_monomial (x : mv_polynomial σ R) (i : σ) : + (x * X i) %ᵐᵒⁿᵒᵐⁱᵃˡ (finsupp.single i 1) = 0 := +mul_monomial_mod_monomial _ _ + +@[simp] lemma mod_monomial_X (i : σ) : + (X i : mv_polynomial σ R) %ᵐᵒⁿᵒᵐⁱᵃˡ (finsupp.single i 1) = 0 := +monomial_mod_monomial _ + +lemma div_monomial_add_mod_monomial_single (x : mv_polynomial σ R) (i : σ) : + X i * (x /ᵐᵒⁿᵒᵐⁱᵃˡ finsupp.single i 1) + x %ᵐᵒⁿᵒᵐⁱᵃˡ finsupp.single i 1 = x := +div_monomial_add_mod_monomial _ _ + +lemma mod_monomial_add_div_monomial_single (x : mv_polynomial σ R) (i : σ) : + x %ᵐᵒⁿᵒᵐⁱᵃˡ finsupp.single i 1 + X i * (x /ᵐᵒⁿᵒᵐⁱᵃˡ finsupp.single i 1) = x := +mod_monomial_add_div_monomial _ _ + +end X_lemmas + +end mv_polynomial diff --git a/src/data/polynomial/inductions.lean b/src/data/polynomial/inductions.lean index 7dd0f627543db..abc02313996b3 100644 --- a/src/data/polynomial/inductions.lean +++ b/src/data/polynomial/inductions.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Damiano Testa. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes, Johannes Hölzl, Scott Morrison, Damiano Testa, Jens Wagemaker -/ +import algebra.monoid_algebra.division import data.nat.interval import data.polynomial.degree.definitions import data.polynomial.induction @@ -31,21 +32,16 @@ variables [semiring R] {p q : R[X]} /-- `div_X p` returns a polynomial `q` such that `q * X + C (p.coeff 0) = p`. It can be used in a semiring where the usual division algorithm is not possible -/ def div_X (p : R[X]) : R[X] := -∑ n in Ico 0 p.nat_degree, monomial n (p.coeff (n + 1)) +⟨add_monoid_algebra.div_of p.to_finsupp 1⟩ @[simp] lemma coeff_div_X : (div_X p).coeff n = p.coeff (n+1) := -begin - simp only [div_X, coeff_monomial, true_and, finset_sum_coeff, not_lt, - mem_Ico, zero_le, finset.sum_ite_eq', ite_eq_left_iff], - intro h, - rw coeff_eq_zero_of_nat_degree_lt (nat.lt_succ_of_le h) -end +by { rw [add_comm], cases p, refl } lemma div_X_mul_X_add (p : R[X]) : div_X p * X + C (p.coeff 0) = p := ext $ by rintro ⟨_|_⟩; simp [coeff_C, nat.succ_ne_zero, coeff_mul_X] @[simp] lemma div_X_C (a : R) : div_X (C a) = 0 := -ext $ λ n, by simp [div_X, coeff_C]; simp [coeff] +ext $ λ n, by simp [coeff_div_X, coeff_C, finsupp.single_eq_of_ne _] lemma div_X_eq_zero_iff : div_X p = 0 ↔ p = C (p.coeff 0) := ⟨λ h, by simpa [eq_comm, h] using div_X_mul_X_add p, From 8ef6f08ff8c781c5c07a8b12843710e1a0d8a688 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Thu, 23 Mar 2023 07:52:55 +0000 Subject: [PATCH 0672/1291] chore(*): add mathlib4 synchronization comments (#18634) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.lie.subalgebra` * `category_theory.limits.shapes.kernel_pair` * `category_theory.monoidal.preadditive` * `category_theory.preadditive.single_obj` * `linear_algebra.sesquilinear_form` * `number_theory.padics.padic_val` * `ring_theory.localization.fraction_ring` * `topology.algebra.uniform_ring` * `topology.covering` * `topology.fiber_bundle.basic` --- src/algebra/lie/subalgebra.lean | 3 +++ src/category_theory/limits/shapes/kernel_pair.lean | 3 +++ src/category_theory/monoidal/preadditive.lean | 3 +++ src/category_theory/preadditive/single_obj.lean | 3 +++ src/linear_algebra/sesquilinear_form.lean | 3 +++ src/number_theory/padics/padic_val.lean | 3 +++ src/ring_theory/localization/fraction_ring.lean | 3 +++ src/topology/algebra/uniform_ring.lean | 3 +++ src/topology/covering.lean | 3 +++ src/topology/fiber_bundle/basic.lean | 3 +++ 10 files changed, 30 insertions(+) diff --git a/src/algebra/lie/subalgebra.lean b/src/algebra/lie/subalgebra.lean index db9b94416b120..7c39d3779ae92 100644 --- a/src/algebra/lie/subalgebra.lean +++ b/src/algebra/lie/subalgebra.lean @@ -9,6 +9,9 @@ import ring_theory.noetherian /-! # Lie subalgebras +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines Lie subalgebras of a Lie algebra and provides basic related definitions and results. diff --git a/src/category_theory/limits/shapes/kernel_pair.lean b/src/category_theory/limits/shapes/kernel_pair.lean index 45483383de7ee..aea26e543051c 100644 --- a/src/category_theory/limits/shapes/kernel_pair.lean +++ b/src/category_theory/limits/shapes/kernel_pair.lean @@ -10,6 +10,9 @@ import category_theory.limits.shapes.regular_mono /-! # Kernel pairs +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines what it means for a parallel pair of morphisms `a b : R ⟶ X` to be the kernel pair for a morphism `f`. Some properties of kernel pairs are given, namely allowing one to transfer between diff --git a/src/category_theory/monoidal/preadditive.lean b/src/category_theory/monoidal/preadditive.lean index 9fb5c2b68234f..8f8ab46005f02 100644 --- a/src/category_theory/monoidal/preadditive.lean +++ b/src/category_theory/monoidal/preadditive.lean @@ -9,6 +9,9 @@ import category_theory.monoidal.functor /-! # Preadditive monoidal categories +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A monoidal category is `monoidal_preadditive` if it is preadditive and tensor product of morphisms is linear in both factors. -/ diff --git a/src/category_theory/preadditive/single_obj.lean b/src/category_theory/preadditive/single_obj.lean index cc81ad832307b..f92d716ddd20a 100644 --- a/src/category_theory/preadditive/single_obj.lean +++ b/src/category_theory/preadditive/single_obj.lean @@ -9,6 +9,9 @@ import category_theory.single_obj /-! # `single_obj α` is preadditive when `α` is a ring. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + -/ namespace category_theory diff --git a/src/linear_algebra/sesquilinear_form.lean b/src/linear_algebra/sesquilinear_form.lean index c7817ccaadfcb..b9e8e550fb951 100644 --- a/src/linear_algebra/sesquilinear_form.lean +++ b/src/linear_algebra/sesquilinear_form.lean @@ -12,6 +12,9 @@ import ring_theory.non_zero_divisors /-! # Sesquilinear form +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This files provides properties about sesquilinear forms. The maps considered are of the form `M₁ →ₛₗ[I₁] M₂ →ₛₗ[I₂] R`, where `I₁ : R₁ →+* R` and `I₂ : R₂ →+* R` are ring homomorphisms and `M₁` is a module over `R₁` and `M₂` is a module over `R₂`. diff --git a/src/number_theory/padics/padic_val.lean b/src/number_theory/padics/padic_val.lean index 67493b8b5f2bf..f6689bb034751 100644 --- a/src/number_theory/padics/padic_val.lean +++ b/src/number_theory/padics/padic_val.lean @@ -10,6 +10,9 @@ import tactic.ring_exp /-! # p-adic Valuation +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the `p`-adic valuation on `ℕ`, `ℤ`, and `ℚ`. The `p`-adic valuation on `ℚ` is the difference of the multiplicities of `p` in the numerator and diff --git a/src/ring_theory/localization/fraction_ring.lean b/src/ring_theory/localization/fraction_ring.lean index d7d0d63870b5a..3f54a8978216d 100644 --- a/src/ring_theory/localization/fraction_ring.lean +++ b/src/ring_theory/localization/fraction_ring.lean @@ -9,6 +9,9 @@ import ring_theory.localization.basic /-! # Fraction ring / fraction field Frac(R) as localization +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions * `is_fraction_ring R K` expresses that `K` is a field of fractions of `R`, as an abbreviation of diff --git a/src/topology/algebra/uniform_ring.lean b/src/topology/algebra/uniform_ring.lean index c2bc74ae60a50..7d8bbd1c1196b 100644 --- a/src/topology/algebra/uniform_ring.lean +++ b/src/topology/algebra/uniform_ring.lean @@ -10,6 +10,9 @@ import topology.algebra.ring.ideal /-! # Completion of topological rings: +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This files endows the completion of a topological ring with a ring structure. More precisely the instance `uniform_space.completion.ring` builds a ring structure on the completion of a ring endowed with a compatible uniform structure in the sense of diff --git a/src/topology/covering.lean b/src/topology/covering.lean index 79f0fbf19cf3f..58f758a2ee4ed 100644 --- a/src/topology/covering.lean +++ b/src/topology/covering.lean @@ -9,6 +9,9 @@ import topology.fiber_bundle.basic /-! # Covering Maps +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines covering maps. ## Main definitions diff --git a/src/topology/fiber_bundle/basic.lean b/src/topology/fiber_bundle/basic.lean index 13900882421a6..48ad8e2aba735 100644 --- a/src/topology/fiber_bundle/basic.lean +++ b/src/topology/fiber_bundle/basic.lean @@ -8,6 +8,9 @@ import topology.fiber_bundle.trivialization /-! # Fiber bundles +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Mathematically, a (topological) fiber bundle with fiber `F` over a base `B` is a space projecting on `B` for which the fibers are all homeomorphic to `F`, such that the local situation around each point is a direct product. From dc9e5ba64653e017743ba5d2c28e42f9f486bf99 Mon Sep 17 00:00:00 2001 From: negiizhao Date: Thu, 23 Mar 2023 10:21:49 +0000 Subject: [PATCH 0673/1291] chore(set_theory/game/pgame): golf le and lf basic API (#18498) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Violeta Hernández [vi.hdz.p@gmail.com](mailto:vi.hdz.p@gmail.com) Co-authored-by: Violeta Hernández Co-authored-by: Eric Wieser --- src/set_theory/game/pgame.lean | 80 ++++++++++++++-------------------- 1 file changed, 32 insertions(+), 48 deletions(-) diff --git a/src/set_theory/game/pgame.lean b/src/set_theory/game/pgame.lean index 15e2be320f042..fd8e873590ce1 100644 --- a/src/set_theory/game/pgame.lean +++ b/src/set_theory/game/pgame.lean @@ -6,6 +6,7 @@ Authors: Reid Barton, Mario Carneiro, Isabel Longbottom, Scott Morrison import data.fin.basic import data.list.basic import logic.relation +import order.game_add /-! # Combinatorial (pre-)games. @@ -255,73 +256,58 @@ instance is_empty_one_right_moves : is_empty (right_moves 1) := pempty.is_empty /-! ### Pre-game order relations -/ -/-- Define simultaneously by mutual induction the `≤` relation and its swapped converse `⧏` on - pre-games. - - The ZFC definition says that `x = {xL | xR}` is less or equal to `y = {yL | yR}` if - `∀ x₁ ∈ xL, x₁ ⧏ y` and `∀ y₂ ∈ yR, x ⧏ y₂`, where `x ⧏ y` means `¬ y ≤ x`. This is a tricky - induction because it only decreases one side at a time, and it also swaps the arguments in the - definition of `≤`. The solution is to define `x ≤ y` and `x ⧏ y` simultaneously. -/ -def le_lf : Π (x y : pgame.{u}), Prop × Prop -| (mk xl xr xL xR) (mk yl yr yL yR) := - -- the orderings of the clauses here are carefully chosen so that - -- and.left/or.inl refer to moves by Left, and - -- and.right/or.inr refer to moves by Right. -((∀ i, (le_lf (xL i) ⟨yl, yr, yL, yR⟩).2) ∧ ∀ j, (le_lf ⟨xl, xr, xL, xR⟩ (yR j)).2, - (∃ i, (le_lf ⟨xl, xr, xL, xR⟩ (yL i)).1) ∨ ∃ j, (le_lf (xR j) ⟨yl, yr, yL, yR⟩).1) -using_well_founded { dec_tac := pgame_wf_tac } - /-- The less or equal relation on pre-games. If `0 ≤ x`, then Left can win `x` as the second player. -/ -instance : has_le pgame := ⟨λ x y, (le_lf x y).1⟩ +instance : has_le pgame := +⟨sym2.game_add.fix wf_is_option $ λ x y le, + (∀ i, ¬ le y (x.move_left i) (sym2.game_add.snd_fst $ is_option.move_left i)) ∧ + (∀ j, ¬ le (y.move_right j) x (sym2.game_add.fst_snd $ is_option.move_right j))⟩ /-- The less or fuzzy relation on pre-games. If `0 ⧏ x`, then Left can win `x` as the first player. -/ -def lf (x y : pgame) : Prop := (le_lf x y).2 +def lf (x y : pgame) : Prop := ¬ y ≤ x localized "infix (name := pgame.lf) ` ⧏ `:50 := pgame.lf" in pgame +@[simp] protected theorem not_le {x y : pgame} : ¬ x ≤ y ↔ y ⧏ x := iff.rfl +@[simp] theorem not_lf {x y : pgame} : ¬ x ⧏ y ↔ y ≤ x := not_not +theorem _root_.has_le.le.not_gf {x y : pgame} : x ≤ y → ¬ y ⧏ x := not_lf.2 +theorem lf.not_ge {x y : pgame} : x ⧏ y → ¬ y ≤ x := id + +/-- Definition of `x ≤ y` on pre-games, in terms of `⧏`. + +The ordering here is chosen so that `and.left` refer to moves by Left, and `and.right` refer to +moves by Right. -/ + +theorem le_iff_forall_lf {x y : pgame} : + x ≤ y ↔ (∀ i, x.move_left i ⧏ y) ∧ ∀ j, x ⧏ y.move_right j := +by { unfold has_le.le, rw sym2.game_add.fix_eq, refl } + /-- Definition of `x ≤ y` on pre-games built using the constructor. -/ @[simp] theorem mk_le_mk {xl xr xL xR yl yr yL yR} : mk xl xr xL xR ≤ mk yl yr yL yR ↔ (∀ i, xL i ⧏ mk yl yr yL yR) ∧ ∀ j, mk xl xr xL xR ⧏ yR j := -show (le_lf _ _).1 ↔ _, by { rw le_lf, refl } - -/-- Definition of `x ≤ y` on pre-games, in terms of `⧏` -/ -theorem le_iff_forall_lf {x y : pgame} : - x ≤ y ↔ (∀ i, x.move_left i ⧏ y) ∧ ∀ j, x ⧏ y.move_right j := -by { cases x, cases y, exact mk_le_mk } +le_iff_forall_lf theorem le_of_forall_lf {x y : pgame} (h₁ : ∀ i, x.move_left i ⧏ y) (h₂ : ∀ j, x ⧏ y.move_right j) : x ≤ y := le_iff_forall_lf.2 ⟨h₁, h₂⟩ -/-- Definition of `x ⧏ y` on pre-games built using the constructor. -/ -@[simp] theorem mk_lf_mk {xl xr xL xR yl yr yL yR} : - mk xl xr xL xR ⧏ mk yl yr yL yR ↔ - (∃ i, mk xl xr xL xR ≤ yL i) ∨ ∃ j, xR j ≤ mk yl yr yL yR := -show (le_lf _ _).2 ↔ _, by { rw le_lf, refl } +/-- Definition of `x ⧏ y` on pre-games, in terms of `≤`. -/-- Definition of `x ⧏ y` on pre-games, in terms of `≤` -/ +The ordering here is chosen so that `or.inl` refer to moves by Left, and `or.inr` refer to +moves by Right. -/ theorem lf_iff_exists_le {x y : pgame} : x ⧏ y ↔ (∃ i, x ≤ y.move_left i) ∨ ∃ j, x.move_right j ≤ y := -by { cases x, cases y, exact mk_lf_mk } - -private theorem not_le_lf {x y : pgame} : (¬ x ≤ y ↔ y ⧏ x) ∧ (¬ x ⧏ y ↔ y ≤ x) := -begin - induction x with xl xr xL xR IHxl IHxr generalizing y, - induction y with yl yr yL yR IHyl IHyr, - simp only [mk_le_mk, mk_lf_mk, IHxl, IHxr, IHyl, IHyr, - not_and_distrib, not_or_distrib, not_forall, not_exists, - and_comm, or_comm, iff_self, and_self] -end +by { rw [lf, le_iff_forall_lf, not_and_distrib], simp } -@[simp] protected theorem not_le {x y : pgame} : ¬ x ≤ y ↔ y ⧏ x := not_le_lf.1 -@[simp] theorem not_lf {x y : pgame} : ¬ x ⧏ y ↔ y ≤ x := not_le_lf.2 -theorem _root_.has_le.le.not_gf {x y : pgame} : x ≤ y → ¬ y ⧏ x := not_lf.2 -theorem lf.not_ge {x y : pgame} : x ⧏ y → ¬ y ≤ x := pgame.not_le.2 +/-- Definition of `x ⧏ y` on pre-games built using the constructor. -/ +@[simp] theorem mk_lf_mk {xl xr xL xR yl yr yL yR} : + mk xl xr xL xR ⧏ mk yl yr yL yR ↔ + (∃ i, mk xl xr xL xR ≤ yL i) ∨ ∃ j, xR j ≤ mk yl yr yL yR := +lf_iff_exists_le theorem le_or_gf (x y : pgame) : x ≤ y ∨ y ⧏ x := by { rw ←pgame.not_le, apply em } @@ -364,8 +350,6 @@ le_of_forall_lf (λ i, pgame.not_le.1 $ λ h, (h₁ hyz h).not_gf $ hxy.move_left_lf i) (λ j, pgame.not_le.1 $ λ h, (h₂ h hxy).not_gf $ hyz.lf_move_right j) -instance : has_lt pgame := ⟨λ x y, x ≤ y ∧ x ⧏ y⟩ - instance : preorder pgame := { le_refl := λ x, begin induction x with _ _ _ _ IHl IHr, @@ -383,8 +367,8 @@ instance : preorder pgame := le_trans_aux (λ i, (IHyl i).2.2) (λ j, (IHxr j).1), le_trans_aux (λ i, (IHzl i).1) (λ j, (IHyr j).2.1)⟩ end, - lt_iff_le_not_le := λ x y, by { rw pgame.not_le, refl }, - ..pgame.has_le, ..pgame.has_lt } + lt := λ x y, x ≤ y ∧ x ⧏ y, + ..pgame.has_le, } theorem lt_iff_le_and_lf {x y : pgame} : x < y ↔ x ≤ y ∧ x ⧏ y := iff.rfl theorem lt_of_le_of_lf {x y : pgame} (h₁ : x ≤ y) (h₂ : x ⧏ y) : x < y := ⟨h₁, h₂⟩ From 3353f661228bd27f632c600cd1a58b874d847c90 Mon Sep 17 00:00:00 2001 From: negiizhao Date: Thu, 23 Mar 2023 15:25:39 +0000 Subject: [PATCH 0674/1291] =?UTF-8?q?chore(*):=20golf=20using=20`acc=5Flif?= =?UTF-8?q?t=E2=82=82=5Fiff`=20and=20`well=5Ffounded=5Flift=E2=82=82=5Fiff?= =?UTF-8?q?`=20(#18526)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Eric Wieser --- src/order/antisymmetrization.lean | 15 +++------ src/order/rel_iso/basic.lean | 54 +++++++++++++++++++++++++------ src/set_theory/zfc/basic.lean | 9 +++--- 3 files changed, 54 insertions(+), 24 deletions(-) diff --git a/src/order/antisymmetrization.lean b/src/order/antisymmetrization.lean index 33d319b56e7d7..fbb41b3be81ad 100644 --- a/src/order/antisymmetrization.lean +++ b/src/order/antisymmetrization.lean @@ -117,13 +117,11 @@ lemma antisymmetrization_fibration : by { rintro a ⟨b⟩ h, exact ⟨b, h, rfl⟩ } lemma acc_antisymmetrization_iff : acc (<) (to_antisymmetrization (≤) a) ↔ acc (<) a := -⟨λ h, by { have := inv_image.accessible _ h, exact this }, - acc.of_fibration _ antisymmetrization_fibration⟩ +acc_lift_on₂'_iff lemma well_founded_antisymmetrization_iff : well_founded (@has_lt.lt (antisymmetrization α (≤)) _) ↔ well_founded (@has_lt.lt α _) := -⟨λ h, ⟨λ a, acc_antisymmetrization_iff.1 $ h.apply _⟩, - λ h, ⟨by { rintro ⟨a⟩, exact acc_antisymmetrization_iff.2 (h.apply a) }⟩⟩ +well_founded_lift_on₂'_iff instance [well_founded_lt α] : well_founded_lt (antisymmetrization α (≤)) := ⟨well_founded_antisymmetrization_iff.2 is_well_founded.wf⟩ @@ -144,13 +142,11 @@ instance [@decidable_rel α (≤)] [@decidable_rel α (<)] [is_total α (≤)] : @[simp] lemma of_antisymmetrization_le_of_antisymmetrization_iff {a b : antisymmetrization α (≤)} : of_antisymmetrization (≤) a ≤ of_antisymmetrization (≤) b ↔ a ≤ b := -by convert to_antisymmetrization_le_to_antisymmetrization_iff.symm; - exact (to_antisymmetrization_of_antisymmetrization _ _).symm +rel_embedding.map_rel_iff (quotient.out'_rel_embedding _) @[simp] lemma of_antisymmetrization_lt_of_antisymmetrization_iff {a b : antisymmetrization α (≤)} : of_antisymmetrization (≤) a < of_antisymmetrization (≤) b ↔ a < b := -by convert to_antisymmetrization_lt_to_antisymmetrization_iff.symm; - exact (to_antisymmetrization_of_antisymmetrization _ _).symm +(quotient.out'_rel_embedding _).map_rel_iff @[mono] lemma to_antisymmetrization_mono : monotone (@to_antisymmetrization α (≤) _) := λ a b, id @@ -184,8 +180,7 @@ variables (α) /-- `of_antisymmetrization` as an order embedding. -/ @[simps] noncomputable def order_embedding.of_antisymmetrization : antisymmetrization α (≤) ↪o α := { to_fun := of_antisymmetrization _, - inj' := λ _ _, quotient.out_inj.1, - map_rel_iff' := λ a b, of_antisymmetrization_le_of_antisymmetrization_iff } + ..quotient.out'_rel_embedding _ } /-- `antisymmetrization` and `order_dual` commute. -/ def order_iso.dual_antisymmetrization : diff --git a/src/order/rel_iso/basic.lean b/src/order/rel_iso/basic.lean index f2bb56d779227..6652b264bfe09 100644 --- a/src/order/rel_iso/basic.lean +++ b/src/order/rel_iso/basic.lean @@ -303,8 +303,14 @@ protected theorem well_founded : ∀ (f : r ↪r s) (h : well_founded s), well_f protected theorem is_well_order : ∀ (f : r ↪r s) [is_well_order β s], is_well_order α r | f H := by exactI {wf := f.well_founded H.wf, ..f.is_strict_total_order} +/-- `quotient.mk` as a relation homomorphism between the relation and the lift of a relation. -/ +@[simps] def _root_.quotient.mk_rel_hom [setoid α] {r : α → α → Prop} (H) : + r →r quotient.lift₂ r H := +⟨@quotient.mk α _, λ _ _, id⟩ + /-- `quotient.out` as a relation embedding between the lift of a relation and the relation. -/ -@[simps] noncomputable def _root_.quotient.out_rel_embedding [s : setoid α] {r : α → α → Prop} (H) : +@[simps] +noncomputable def _root_.quotient.out_rel_embedding [setoid α] {r : α → α → Prop} (H) : quotient.lift₂ r H ↪r r := ⟨embedding.quotient_out α, begin refine λ x y, quotient.induction_on₂ x y (λ a b, _), @@ -312,20 +318,50 @@ protected theorem is_well_order : ∀ (f : r ↪r s) [is_well_order β s], is_we apply quotient.mk_out end⟩ +/-- `quotient.out'` as a relation embedding between the lift of a relation and the relation. -/ +@[simps] +noncomputable def _root_.quotient.out'_rel_embedding {s : setoid α} {r : α → α → Prop} (H) : + (λ a b, quotient.lift_on₂' a b r H) ↪r r := +{ to_fun := quotient.out', + ..quotient.out_rel_embedding _ } + +@[simp] theorem _root_.acc_lift₂_iff [setoid α] {r : α → α → Prop} {H} {a} : + acc (quotient.lift₂ r H) ⟦a⟧ ↔ acc r a := +begin + split, + { exact rel_hom_class.acc (quotient.mk_rel_hom H) a, }, + { intro ac, + induction ac with _ H IH, dsimp at IH, + refine ⟨_, λ q h, _⟩, + obtain ⟨a', rfl⟩ := q.exists_rep, + exact IH a' h, }, +end + +@[simp] theorem _root_.acc_lift_on₂'_iff {s : setoid α} {r : α → α → Prop} {H} {a} : + acc (λ x y, quotient.lift_on₂' x y r H) (quotient.mk' a : quotient s) ↔ acc r a := +acc_lift₂_iff + /-- A relation is well founded iff its lift to a quotient is. -/ -@[simp] theorem _root_.well_founded_lift₂_iff [s : setoid α] {r : α → α → Prop} {H} : +theorem _root_.well_founded_lift₂_iff [setoid α] {r : α → α → Prop} {H} : well_founded (quotient.lift₂ r H) ↔ well_founded r := -⟨λ hr, begin - suffices : ∀ {x : quotient s} {a : α}, ⟦a⟧ = x → acc r a, - { exact ⟨λ a, this rfl⟩ }, - { refine λ x, hr.induction x _, - rintros x IH a rfl, - exact ⟨_, λ b hb, IH ⟦b⟧ hb rfl⟩ } -end, (quotient.out_rel_embedding H).well_founded⟩ +begin + split, + { exact rel_hom_class.well_founded (quotient.mk_rel_hom H), }, + { refine λ wf, ⟨λ q, _⟩, + obtain ⟨a, rfl⟩ := q.exists_rep, + exact acc_lift₂_iff.2 (wf.apply a), }, +end alias _root_.well_founded_lift₂_iff ↔ _root_.well_founded.of_quotient_lift₂ _root_.well_founded.quotient_lift₂ +@[simp] theorem _root_.well_founded_lift_on₂'_iff {s : setoid α} {r : α → α → Prop} {H} : + well_founded (λ x y : quotient s, quotient.lift_on₂' x y r H) ↔ well_founded r := +well_founded_lift₂_iff + +alias _root_.well_founded_lift_on₂'_iff ↔ + _root_.well_founded.of_quotient_lift_on₂' _root_.well_founded.quotient_lift_on₂' + /-- To define an relation embedding from an antisymmetric relation `r` to a reflexive relation `s` it suffices to give a function together with a proof that it satisfies `s (f a) (f b) ↔ r a b`. diff --git a/src/set_theory/zfc/basic.lean b/src/set_theory/zfc/basic.lean index 06dbfb121373a..5aa704b4fa60d 100644 --- a/src/set_theory/zfc/basic.lean +++ b/src/set_theory/zfc/basic.lean @@ -720,14 +720,13 @@ by { rw ←mem_to_set, simp } @[simp] theorem mem_diff {x y z : Set.{u}} : z ∈ x \ y ↔ z ∈ x ∧ z ∉ y := @@mem_sep (λ z : Set.{u}, z ∉ y) +theorem mem_wf : @well_founded Set (∈) := +well_founded_lift₂_iff.mpr pSet.mem_wf + /-- Induction on the `∈` relation. -/ @[elab_as_eliminator] theorem induction_on {p : Set → Prop} (x) (h : ∀ x, (∀ y ∈ x, p y) → p x) : p x := -quotient.induction_on x $ λ u, pSet.rec_on u $ λ α A IH, h _ $ λ y, -show @has_mem.mem _ _ Set.has_mem y ⟦⟨α, A⟩⟧ → p y, from -quotient.induction_on y (λ v ⟨a, ha⟩, by { rw (@quotient.sound pSet _ _ _ ha), exact IH a }) - -theorem mem_wf : @well_founded Set (∈) := ⟨λ x, induction_on x acc.intro⟩ +mem_wf.induction x h instance : has_well_founded Set := ⟨_, mem_wf⟩ From b19481deb571022990f1baa9cbf9172e6757a479 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 23 Mar 2023 18:37:37 +0000 Subject: [PATCH 0675/1291] feat(order/antichain): Antichains are order-connected (#18636) --- src/data/set/intervals/ord_connected.lean | 10 ++++++++++ src/order/antichain.lean | 13 ++++++++++++- 2 files changed, 22 insertions(+), 1 deletion(-) diff --git a/src/data/set/intervals/ord_connected.lean b/src/data/set/intervals/ord_connected.lean index 7dd30593bff4a..2a7537495a4e3 100644 --- a/src/data/set/intervals/ord_connected.lean +++ b/src/data/set/intervals/ord_connected.lean @@ -5,6 +5,7 @@ Authors: Yury G. Kudryashov -/ import data.set.intervals.unordered_interval import data.set.lattice +import order.antichain /-! # Order-connected sets @@ -173,6 +174,15 @@ dual_ord_connected_iff.2 ‹_› end preorder +section partial_order +variables {α : Type*} [partial_order α] {s : set α} + +protected lemma is_antichain.ord_connected (hs : is_antichain (≤) s) : s.ord_connected := +⟨λ x hx y hy z hz, by { obtain rfl := hs.eq hx hy (hz.1.trans hz.2), + rw [Icc_self, mem_singleton_iff] at hz, rwa hz }⟩ + +end partial_order + section linear_order variables {α : Type*} [linear_order α] {s : set α} {x : α} diff --git a/src/order/antichain.lean b/src/order/antichain.lean index ab3d0e5e6bbc0..a49fe4cf4beeb 100644 --- a/src/order/antichain.lean +++ b/src/order/antichain.lean @@ -26,7 +26,7 @@ relation is `G.adj` for `G : simple_graph α`, this corresponds to independent s open function set section general -variables {α β : Type*} {r r₁ r₂ : α → α → Prop} {r' : β → β → Prop} {s t : set α} {a : α} +variables {α β : Type*} {r r₁ r₂ : α → α → Prop} {r' : β → β → Prop} {s t : set α} {a b : α} protected lemma symmetric.compl (h : symmetric r) : symmetric rᶜ := λ x y hr hr', hr $ h hr' @@ -176,6 +176,9 @@ hs.pairwise _ section preorder variables [preorder α] +lemma is_antichain.not_lt (hs : is_antichain (≤) s) (ha : a ∈ s) (hb : b ∈ s) : ¬ a < b := +λ h, hs ha hb h.ne h.le + lemma is_antichain_and_least_iff : is_antichain (≤) s ∧ is_least s a ↔ s = {a} := ⟨λ h, eq_singleton_iff_unique_mem.2 ⟨h.2.1, λ b hb, h.1.eq' hb h.2.1 (h.2.2 hb)⟩, by { rintro rfl, exact ⟨is_antichain_singleton _ _, is_least_singleton⟩ }⟩ @@ -204,6 +207,14 @@ is_greatest_top_iff.symm.trans hs.greatest_iff end preorder +section partial_order +variables [partial_order α] + +lemma is_antichain_iff_forall_not_lt : is_antichain (≤) s ↔ ∀ ⦃a⦄, a ∈ s → ∀ ⦃b⦄, b ∈ s → ¬ a < b := +⟨λ hs a ha b, hs.not_lt ha, λ hs a ha b hb h h', hs ha hb $ h'.lt_of_ne h⟩ + +end partial_order + /-! ### Strong antichains -/ /-- A strong (upward) antichain is a set such that no two distinct elements are related to a common From 842557b6253ace82c125d567a80b5f74f6ce9f99 Mon Sep 17 00:00:00 2001 From: Michael Stoll Date: Thu, 23 Mar 2023 21:18:31 +0000 Subject: [PATCH 0676/1291] chore(number_theory/pell): Move def. of solution type and API to beginning (#18639) On the suggestion of @eric-wieser, this PR separates moving the definition of the type of solutions and the API lemmas to the beginning of the file, to help with the reviewing of #18626. Co-authored-by: Eric Wieser --- src/number_theory/pell.lean | 197 ++++++++++++++++++------------------ 1 file changed, 100 insertions(+), 97 deletions(-) diff --git a/src/number_theory/pell.lean b/src/number_theory/pell.lean index 656a1950447e8..f6fb65d95dae1 100644 --- a/src/number_theory/pell.lean +++ b/src/number_theory/pell.lean @@ -41,103 +41,6 @@ Pell's equation namespace pell -section existence - -variables {d : ℤ} - -open set real - -/-- If `d` is a positive integer that is not a square, then there is a nontrivial solution -to the Pell equation `x^2 - d*y^2 = 1`. -/ -theorem exists_of_not_is_square (h₀ : 0 < d) (hd : ¬ is_square d) : - ∃ x y : ℤ, x ^ 2 - d * y ^ 2 = 1 ∧ y ≠ 0 := -begin - let ξ : ℝ := sqrt d, - have hξ : irrational ξ, - { refine irrational_nrt_of_notint_nrt 2 d (sq_sqrt $ int.cast_nonneg.mpr h₀.le) _ two_pos, - rintro ⟨x, hx⟩, - refine hd ⟨x, @int.cast_injective ℝ _ _ d (x * x) _⟩, - rw [← sq_sqrt $ int.cast_nonneg.mpr h₀.le, int.cast_mul, ← hx, sq], }, - obtain ⟨M, hM₁⟩ := exists_int_gt (2 * |ξ| + 1), - have hM : {q : ℚ | |q.1 ^ 2 - d * q.2 ^ 2| < M}.infinite, - { refine infinite.mono (λ q h, _) (infinite_rat_abs_sub_lt_one_div_denom_sq_of_irrational hξ), - have h0 : 0 < (q.2 : ℝ) ^ 2 := pow_pos (nat.cast_pos.mpr q.pos) 2, - have h1 : (q.num : ℝ) / (q.denom : ℝ) = q := by exact_mod_cast q.num_div_denom, - rw [mem_set_of, abs_sub_comm, ← @int.cast_lt ℝ, ← div_lt_div_right (abs_pos_of_pos h0)], - push_cast, - rw [← abs_div, abs_sq, sub_div, mul_div_cancel _ h0.ne', - ← div_pow, h1, ← sq_sqrt (int.cast_pos.mpr h₀).le, sq_sub_sq, abs_mul, ← mul_one_div], - refine mul_lt_mul'' (((abs_add ξ q).trans _).trans_lt hM₁) h (abs_nonneg _) (abs_nonneg _), - rw [two_mul, add_assoc, add_le_add_iff_left, ← sub_le_iff_le_add'], - rw [mem_set_of, abs_sub_comm] at h, - refine (abs_sub_abs_le_abs_sub (q : ℝ) ξ).trans (h.le.trans _), - rw [div_le_one h0, one_le_sq_iff_one_le_abs, nat.abs_cast, nat.one_le_cast], - exact q.pos, }, - obtain ⟨m, hm⟩ : ∃ m : ℤ, {q : ℚ | q.1 ^ 2 - d * q.2 ^ 2 = m}.infinite, - { contrapose! hM, - simp only [not_infinite] at hM ⊢, - refine (congr_arg _ (ext (λ x, _))).mp (finite.bUnion (finite_Ioo (-M) M) (λ m _, hM m)), - simp only [abs_lt, mem_set_of_eq, mem_Ioo, mem_Union, exists_prop, exists_eq_right'], }, - have hm₀ : m ≠ 0, - { rintro rfl, - obtain ⟨q, hq⟩ := hm.nonempty, - rw [mem_set_of, sub_eq_zero, mul_comm] at hq, - obtain ⟨a, ha⟩ := (int.pow_dvd_pow_iff two_pos).mp ⟨d, hq⟩, - rw [ha, mul_pow, mul_right_inj' (pow_pos (int.coe_nat_pos.mpr q.pos) 2).ne'] at hq, - exact hd ⟨a, sq a ▸ hq.symm⟩, }, - haveI := ne_zero_iff.mpr (int.nat_abs_ne_zero.mpr hm₀), - let f : ℚ → (zmod m.nat_abs) × (zmod m.nat_abs) := λ q, (q.1, q.2), - obtain ⟨q₁, h₁ : q₁.1 ^ 2 - d * q₁.2 ^ 2 = m, q₂, h₂ : q₂.1 ^ 2 - d * q₂.2 ^ 2 = m, hne, hqf⟩ := - hm.exists_ne_map_eq_of_maps_to (maps_to_univ f _) finite_univ, - obtain ⟨hq1 : (q₁.1 : zmod m.nat_abs) = q₂.1, hq2 : (q₁.2 : zmod m.nat_abs) = q₂.2⟩ := - prod.ext_iff.mp hqf, - have hd₁ : m ∣ q₁.1 * q₂.1 - d * (q₁.2 * q₂.2), - { rw [← int.nat_abs_dvd, ← zmod.int_coe_zmod_eq_zero_iff_dvd], - push_cast, - rw [hq1, hq2, ← sq, ← sq], - norm_cast, - rw [zmod.int_coe_zmod_eq_zero_iff_dvd, int.nat_abs_dvd, nat.cast_pow, ← h₂], }, - have hd₂ : m ∣ q₁.1 * q₂.2 - q₂.1 * q₁.2, - { rw [← int.nat_abs_dvd, ← zmod.int_coe_eq_int_coe_iff_dvd_sub], - push_cast, - rw [hq1, hq2], }, - replace hm₀ : (m : ℚ) ≠ 0 := int.cast_ne_zero.mpr hm₀, - refine ⟨(q₁.1 * q₂.1 - d * (q₁.2 * q₂.2)) / m, (q₁.1 * q₂.2 - q₂.1 * q₁.2) / m, _, _⟩, - { qify [hd₁, hd₂], - field_simp [hm₀], - norm_cast, - conv_rhs {congr, rw sq, congr, rw ← h₁, skip, rw ← h₂}, - push_cast, - ring, }, - { qify [hd₂], - refine div_ne_zero_iff.mpr ⟨_, hm₀⟩, - exact_mod_cast mt sub_eq_zero.mp (mt rat.eq_iff_mul_eq_mul.mpr hne), }, -end - -/-- If `d` is a positive integer, then there is a nontrivial solution -to the Pell equation `x^2 - d*y^2 = 1` if and only if `d` is not a square. -/ -theorem exists_iff_not_is_square (h₀ : 0 < d) : - (∃ x y : ℤ, x ^ 2 - d * y ^ 2 = 1 ∧ y ≠ 0) ↔ ¬ is_square d := -begin - refine ⟨_, exists_of_not_is_square h₀⟩, - rintros ⟨x, y, hxy, hy⟩ ⟨a, rfl⟩, - rw [← sq, ← mul_pow, sq_sub_sq] at hxy, - simpa [mul_self_pos.mp h₀, sub_eq_add_neg, eq_neg_self_iff] using int.eq_of_mul_eq_one hxy, -end - -/-- If `d` is a positive integer that is not a square, then there exists a solution -to the Pell equation `x^2 - d*y^2 = 1` with `x > 1` and `y > 0`. -/ -lemma exists_pos_of_not_is_square (h₀ : 0 < d) (hd : ¬ is_square d) : - ∃ x y : ℤ, x ^ 2 - d * y ^ 2 = 1 ∧ 1 < x ∧ 0 < y := -begin - obtain ⟨x, y, h, hy⟩ := exists_of_not_is_square h₀ hd, - refine ⟨|x|, |y|, by rwa [sq_abs, sq_abs], _, abs_pos.mpr hy⟩, - rw [← one_lt_sq_iff_one_lt_abs, eq_add_of_sub_eq h, lt_add_iff_pos_right], - exact mul_pos h₀ (sq_pos_of_ne_zero y hy), -end - -end existence - /-! ### Group structure of the solution set @@ -243,4 +146,104 @@ lemma y_neg (a : solution₁ d) : (-a).y = -a.y := rfl end solution₁ +section existence + +/-! +### Existence of nontrivial solutions +-/ + +variables {d : ℤ} + +open set real + +/-- If `d` is a positive integer that is not a square, then there is a nontrivial solution +to the Pell equation `x^2 - d*y^2 = 1`. -/ +theorem exists_of_not_is_square (h₀ : 0 < d) (hd : ¬ is_square d) : + ∃ x y : ℤ, x ^ 2 - d * y ^ 2 = 1 ∧ y ≠ 0 := +begin + let ξ : ℝ := sqrt d, + have hξ : irrational ξ, + { refine irrational_nrt_of_notint_nrt 2 d (sq_sqrt $ int.cast_nonneg.mpr h₀.le) _ two_pos, + rintro ⟨x, hx⟩, + refine hd ⟨x, @int.cast_injective ℝ _ _ d (x * x) _⟩, + rw [← sq_sqrt $ int.cast_nonneg.mpr h₀.le, int.cast_mul, ← hx, sq], }, + obtain ⟨M, hM₁⟩ := exists_int_gt (2 * |ξ| + 1), + have hM : {q : ℚ | |q.1 ^ 2 - d * q.2 ^ 2| < M}.infinite, + { refine infinite.mono (λ q h, _) (infinite_rat_abs_sub_lt_one_div_denom_sq_of_irrational hξ), + have h0 : 0 < (q.2 : ℝ) ^ 2 := pow_pos (nat.cast_pos.mpr q.pos) 2, + have h1 : (q.num : ℝ) / (q.denom : ℝ) = q := by exact_mod_cast q.num_div_denom, + rw [mem_set_of, abs_sub_comm, ← @int.cast_lt ℝ, ← div_lt_div_right (abs_pos_of_pos h0)], + push_cast, + rw [← abs_div, abs_sq, sub_div, mul_div_cancel _ h0.ne', + ← div_pow, h1, ← sq_sqrt (int.cast_pos.mpr h₀).le, sq_sub_sq, abs_mul, ← mul_one_div], + refine mul_lt_mul'' (((abs_add ξ q).trans _).trans_lt hM₁) h (abs_nonneg _) (abs_nonneg _), + rw [two_mul, add_assoc, add_le_add_iff_left, ← sub_le_iff_le_add'], + rw [mem_set_of, abs_sub_comm] at h, + refine (abs_sub_abs_le_abs_sub (q : ℝ) ξ).trans (h.le.trans _), + rw [div_le_one h0, one_le_sq_iff_one_le_abs, nat.abs_cast, nat.one_le_cast], + exact q.pos, }, + obtain ⟨m, hm⟩ : ∃ m : ℤ, {q : ℚ | q.1 ^ 2 - d * q.2 ^ 2 = m}.infinite, + { contrapose! hM, + simp only [not_infinite] at hM ⊢, + refine (congr_arg _ (ext (λ x, _))).mp (finite.bUnion (finite_Ioo (-M) M) (λ m _, hM m)), + simp only [abs_lt, mem_set_of_eq, mem_Ioo, mem_Union, exists_prop, exists_eq_right'], }, + have hm₀ : m ≠ 0, + { rintro rfl, + obtain ⟨q, hq⟩ := hm.nonempty, + rw [mem_set_of, sub_eq_zero, mul_comm] at hq, + obtain ⟨a, ha⟩ := (int.pow_dvd_pow_iff two_pos).mp ⟨d, hq⟩, + rw [ha, mul_pow, mul_right_inj' (pow_pos (int.coe_nat_pos.mpr q.pos) 2).ne'] at hq, + exact hd ⟨a, sq a ▸ hq.symm⟩, }, + haveI := ne_zero_iff.mpr (int.nat_abs_ne_zero.mpr hm₀), + let f : ℚ → (zmod m.nat_abs) × (zmod m.nat_abs) := λ q, (q.1, q.2), + obtain ⟨q₁, h₁ : q₁.1 ^ 2 - d * q₁.2 ^ 2 = m, q₂, h₂ : q₂.1 ^ 2 - d * q₂.2 ^ 2 = m, hne, hqf⟩ := + hm.exists_ne_map_eq_of_maps_to (maps_to_univ f _) finite_univ, + obtain ⟨hq1 : (q₁.1 : zmod m.nat_abs) = q₂.1, hq2 : (q₁.2 : zmod m.nat_abs) = q₂.2⟩ := + prod.ext_iff.mp hqf, + have hd₁ : m ∣ q₁.1 * q₂.1 - d * (q₁.2 * q₂.2), + { rw [← int.nat_abs_dvd, ← zmod.int_coe_zmod_eq_zero_iff_dvd], + push_cast, + rw [hq1, hq2, ← sq, ← sq], + norm_cast, + rw [zmod.int_coe_zmod_eq_zero_iff_dvd, int.nat_abs_dvd, nat.cast_pow, ← h₂], }, + have hd₂ : m ∣ q₁.1 * q₂.2 - q₂.1 * q₁.2, + { rw [← int.nat_abs_dvd, ← zmod.int_coe_eq_int_coe_iff_dvd_sub], + push_cast, + rw [hq1, hq2], }, + replace hm₀ : (m : ℚ) ≠ 0 := int.cast_ne_zero.mpr hm₀, + refine ⟨(q₁.1 * q₂.1 - d * (q₁.2 * q₂.2)) / m, (q₁.1 * q₂.2 - q₂.1 * q₁.2) / m, _, _⟩, + { qify [hd₁, hd₂], + field_simp [hm₀], + norm_cast, + conv_rhs {congr, rw sq, congr, rw ← h₁, skip, rw ← h₂}, + push_cast, + ring, }, + { qify [hd₂], + refine div_ne_zero_iff.mpr ⟨_, hm₀⟩, + exact_mod_cast mt sub_eq_zero.mp (mt rat.eq_iff_mul_eq_mul.mpr hne), }, +end + +/-- If `d` is a positive integer, then there is a nontrivial solution +to the Pell equation `x^2 - d*y^2 = 1` if and only if `d` is not a square. -/ +theorem exists_iff_not_is_square (h₀ : 0 < d) : + (∃ x y : ℤ, x ^ 2 - d * y ^ 2 = 1 ∧ y ≠ 0) ↔ ¬ is_square d := +begin + refine ⟨_, exists_of_not_is_square h₀⟩, + rintros ⟨x, y, hxy, hy⟩ ⟨a, rfl⟩, + rw [← sq, ← mul_pow, sq_sub_sq] at hxy, + simpa [mul_self_pos.mp h₀, sub_eq_add_neg, eq_neg_self_iff] using int.eq_of_mul_eq_one hxy, +end + +/-- If `d` is a positive integer that is not a square, then there exists a solution +to the Pell equation `x^2 - d*y^2 = 1` with `x > 1` and `y > 0`. -/ +lemma exists_pos_of_not_is_square (h₀ : 0 < d) (hd : ¬ is_square d) : + ∃ x y : ℤ, x ^ 2 - d * y ^ 2 = 1 ∧ 1 < x ∧ 0 < y := +begin + obtain ⟨x, y, h, hy⟩ := exists_of_not_is_square h₀ hd, + refine ⟨|x|, |y|, by rwa [sq_abs, sq_abs], _, abs_pos.mpr hy⟩, + rw [← one_lt_sq_iff_one_lt_abs, eq_add_of_sub_eq h, lt_add_iff_pos_right], + exact mul_pos h₀ (sq_pos_of_ne_zero y hy), +end + +end existence end pell From 0caf3701139ef2e69c215717665361cda205a90b Mon Sep 17 00:00:00 2001 From: Amelia Livingston <101damnations@github.com> Date: Thu, 23 Mar 2023 22:20:53 +0000 Subject: [PATCH 0677/1291] feat(representation_theory/Rep): describe monoidal closed structure (#18148) The monoidal closed structure on `Rep k G` defines an internal hom of representations; we show this agrees with `representation.lin_hom`. Moreover, the maps defining the hom-set bijection come from the tensor-hom adjunction for $k$-modules. --- src/algebra/category/Module/monoidal.lean | 26 ++++- src/algebra/category/Mon/basic.lean | 2 + .../closed/functor_category.lean | 9 ++ src/category_theory/closed/monoidal.lean | 26 +++++ src/representation_theory/Action.lean | 59 +++++++++++ src/representation_theory/Rep.lean | 97 ++++++++++++++++++- 6 files changed, 216 insertions(+), 3 deletions(-) diff --git a/src/algebra/category/Module/monoidal.lean b/src/algebra/category/Module/monoidal.lean index 0f80c5596f23f..20dd4ea2f8fc4 100644 --- a/src/algebra/category/Module/monoidal.lean +++ b/src/algebra/category/Module/monoidal.lean @@ -307,6 +307,9 @@ instance : monoidal_closed (Module.{u} R) := adj := adjunction.mk_of_hom_equiv { hom_equiv := λ N P, monoidal_closed_hom_equiv M N P, } } } } +lemma ihom_map_apply {M N P : Module.{u} R} (f : N ⟶ P) (g : Module.of R (M ⟶ N)) : + (ihom M).map f g = g ≫ f := rfl + -- I can't seem to express the function coercion here without writing `@coe_fn`. @[simp] lemma monoidal_closed_curry {M N P : Module.{u} R} (f : M ⊗ N ⟶ P) (x : M) (y : N) : @@ -318,6 +321,27 @@ rfl lemma monoidal_closed_uncurry {M N P : Module.{u} R} (f : N ⟶ (M ⟶[Module.{u} R] P)) (x : M) (y : N) : monoidal_closed.uncurry f (x ⊗ₜ[R] y) = (@coe_fn _ _ linear_map.has_coe_to_fun (f y)) x := -by { simp only [monoidal_closed.uncurry, ihom.adjunction, is_left_adjoint.adj], simp, } +rfl + +/-- Describes the counit of the adjunction `M ⊗ - ⊣ Hom(M, -)`. Given an `R`-module `N` this +should give a map `M ⊗ Hom(M, N) ⟶ N`, so we flip the order of the arguments in the identity map +`Hom(M, N) ⟶ (M ⟶ N)` and uncurry the resulting map `M ⟶ Hom(M, N) ⟶ N.` -/ +lemma ihom_ev_app (M N : Module.{u} R) : + (ihom.ev M).app N = tensor_product.uncurry _ _ _ _ linear_map.id.flip := +begin + ext, + exact Module.monoidal_closed_uncurry _ _ _, +end + +/-- Describes the unit of the adjunction `M ⊗ - ⊣ Hom(M, -)`. Given an `R`-module `N` this should +define a map `N ⟶ Hom(M, M ⊗ N)`, which is given by flipping the arguments in the natural +`R`-bilinear map `M ⟶ N ⟶ M ⊗ N`. -/ +lemma ihom_coev_app (M N : Module.{u} R) : + (ihom.coev M).app N = (tensor_product.mk _ _ _).flip := +rfl + +lemma monoidal_closed_pre_app {M N : Module.{u} R} (P : Module.{u} R) (f : N ⟶ M) : + (monoidal_closed.pre f).app P = linear_map.lcomp R _ f := +rfl end Module diff --git a/src/algebra/category/Mon/basic.lean b/src/algebra/category/Mon/basic.lean index 5b050d6a7f05f..33350da6a4ec4 100644 --- a/src/algebra/category/Mon/basic.lean +++ b/src/algebra/category/Mon/basic.lean @@ -80,6 +80,8 @@ instance (M : Mon) : monoid M := M.str @[simp, to_additive] lemma coe_of (R : Type u) [monoid R] : (Mon.of R : Type u) = R := rfl +@[to_additive] instance {G : Type*} [group G] : group (Mon.of G) := by assumption + end Mon /-- The category of commutative monoids and monoid morphisms. -/ diff --git a/src/category_theory/closed/functor_category.lean b/src/category_theory/closed/functor_category.lean index 079a52c0f2982..c6633e2f078aa 100644 --- a/src/category_theory/closed/functor_category.lean +++ b/src/category_theory/closed/functor_category.lean @@ -70,4 +70,13 @@ with the pointwise monoidal structure, is monoidal closed. -/ @[simps] instance monoidal_closed : monoidal_closed (D ⥤ C) := { closed' := by apply_instance } +lemma ihom_map (F : D ⥤ C) {G H : D ⥤ C} (f : G ⟶ H) : + (ihom F).map f = (closed_ihom F).map f := rfl + +lemma ihom_ev_app (F G : D ⥤ C) : + (ihom.ev F).app G = (closed_counit F).app G := rfl + +lemma ihom_coev_app (F G : D ⥤ C) : + (ihom.coev F).app G = (closed_unit F).app G := rfl + end category_theory.functor diff --git a/src/category_theory/closed/monoidal.lean b/src/category_theory/closed/monoidal.lean index 57f1979923cfb..5016a124b2b32 100644 --- a/src/category_theory/closed/monoidal.lean +++ b/src/category_theory/closed/monoidal.lean @@ -264,6 +264,32 @@ def of_equiv (F : monoidal_functor C D) [is_equivalence F.to_functor] [h : monoi exact adjunction.left_adjoint_of_nat_iso i, end } } +/-- Suppose we have a monoidal equivalence `F : C ≌ D`, with `D` monoidal closed. We can pull the +monoidal closed instance back along the equivalence. For `X, Y, Z : C`, this lemma describes the +resulting currying map `Hom(X ⊗ Y, Z) → Hom(Y, (X ⟶[C] Z))`. (`X ⟶[C] Z` is defined to be +`F⁻¹(F(X) ⟶[D] F(Z))`, so currying in `C` is given by essentially conjugating currying in +`D` by `F.`) -/ +lemma of_equiv_curry_def (F : monoidal_functor C D) [is_equivalence F.to_functor] + [h : monoidal_closed D] {X Y Z : C} (f : X ⊗ Y ⟶ Z) : + @monoidal_closed.curry _ _ _ _ _ _ ((monoidal_closed.of_equiv F).1 _) f = + (F.1.1.adjunction.hom_equiv Y ((ihom _).obj _)) (monoidal_closed.curry + (F.1.1.inv.adjunction.hom_equiv (F.1.1.obj X ⊗ F.1.1.obj Y) Z + ((comp_inv_iso (F.comm_tensor_left X)).hom.app Y ≫ f))) := rfl + +/-- Suppose we have a monoidal equivalence `F : C ≌ D`, with `D` monoidal closed. We can pull the +monoidal closed instance back along the equivalence. For `X, Y, Z : C`, this lemma describes the +resulting uncurrying map `Hom(Y, (X ⟶[C] Z)) → Hom(X ⊗ Y ⟶ Z)`. (`X ⟶[C] Z` is +defined to be `F⁻¹(F(X) ⟶[D] F(Z))`, so uncurrying in `C` is given by essentially conjugating +uncurrying in `D` by `F.`) -/ +lemma of_equiv_uncurry_def + (F : monoidal_functor C D) [is_equivalence F.to_functor] [h : monoidal_closed D] {X Y Z : C} + (f : Y ⟶ (@ihom _ _ _ X $ (monoidal_closed.of_equiv F).1 X).obj Z) : + @monoidal_closed.uncurry _ _ _ _ _ _ ((monoidal_closed.of_equiv F).1 _) f = + (comp_inv_iso (F.comm_tensor_left X)).inv.app Y ≫ (F.1.1.inv.adjunction.hom_equiv + (F.1.1.obj X ⊗ F.1.1.obj Y) Z).symm (monoidal_closed.uncurry + ((F.1.1.adjunction.hom_equiv Y ((ihom (F.1.1.obj X)).obj (F.1.1.obj Z))).symm f)) := +rfl + end of_equiv end monoidal_closed diff --git a/src/representation_theory/Action.lean b/src/representation_theory/Action.lean index 96a4e0500f76f..e20119ba8658c 100644 --- a/src/representation_theory/Action.lean +++ b/src/representation_theory/Action.lean @@ -206,6 +206,12 @@ def functor_category_equivalence : Action V G ≌ (single_obj G ⥤ V) := attribute [simps] functor_category_equivalence +lemma functor_category_equivalence.functor_def : + (functor_category_equivalence V G).functor = functor_category_equivalence.functor := rfl + +lemma functor_category_equivalence.inverse_def : + (functor_category_equivalence V G).inverse = functor_category_equivalence.inverse := rfl + instance [has_finite_products V] : has_finite_products (Action V G) := { out := λ n, adjunction.has_limits_of_shape_of_equivalence (Action.functor_category_equivalence _ _).functor } @@ -453,6 +459,59 @@ 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, } +@[simp] lemma functor_category_monoidal_equivalence.μ_app (A B : Action V G) : + ((functor_category_monoidal_equivalence V G).μ A B).app punit.star = 𝟙 _ := +begin + dunfold functor_category_monoidal_equivalence, + simp only [monoidal.from_transported_to_lax_monoidal_functor_μ], + show (𝟙 A.V ⊗ 𝟙 B.V) ≫ 𝟙 (A.V ⊗ B.V) ≫ (𝟙 A.V ⊗ 𝟙 B.V) = 𝟙 (A.V ⊗ B.V), + simp only [monoidal_category.tensor_id, category.comp_id], +end + +@[simp] lemma functor_category_monoidal_equivalence.μ_iso_inv_app (A B : Action V G) : + ((functor_category_monoidal_equivalence V G).μ_iso A B).inv.app punit.star = 𝟙 _ := +begin + rw [←nat_iso.app_inv, ←is_iso.iso.inv_hom], + refine is_iso.inv_eq_of_hom_inv_id _, + rw [category.comp_id, nat_iso.app_hom, monoidal_functor.μ_iso_hom, + functor_category_monoidal_equivalence.μ_app], +end + +@[simp] lemma functor_category_monoidal_equivalence.ε_app : + (functor_category_monoidal_equivalence V G).ε.app punit.star = 𝟙 _ := +begin + dunfold functor_category_monoidal_equivalence, + simp only [monoidal.from_transported_to_lax_monoidal_functor_ε], + show 𝟙 (monoidal_category.tensor_unit V) ≫ _ = 𝟙 (monoidal_category.tensor_unit V), + rw [nat_iso.is_iso_inv_app, category.id_comp], + exact is_iso.inv_id, +end + +@[simp] lemma functor_category_monoidal_equivalence.inv_counit_app_hom (A : Action V G) : + ((functor_category_monoidal_equivalence _ _).inv.adjunction.counit.app A).hom = 𝟙 _ := +rfl + +@[simp] lemma functor_category_monoidal_equivalence.counit_app (A : single_obj G ⥤ V) : + ((functor_category_monoidal_equivalence _ _).adjunction.counit.app A).app punit.star = 𝟙 _ := rfl + +@[simp] lemma functor_category_monoidal_equivalence.inv_unit_app_app + (A : single_obj G ⥤ V) : + ((functor_category_monoidal_equivalence _ _).inv.adjunction.unit.app A).app + punit.star = 𝟙 _ := rfl + +@[simp] lemma functor_category_monoidal_equivalence.unit_app_hom (A : Action V G) : + ((functor_category_monoidal_equivalence _ _).adjunction.unit.app A).hom = 𝟙 _ := +rfl + +@[simp] lemma functor_category_monoidal_equivalence.functor_map {A B : Action V G} (f : A ⟶ B) : + (functor_category_monoidal_equivalence _ _).1.1.map f + = functor_category_equivalence.functor.map f := rfl + +@[simp] lemma functor_category_monoidal_equivalence.inverse_map + {A B : single_obj G ⥤ V} (f : A ⟶ B) : + (functor_category_monoidal_equivalence _ _).1.inv.map f + = functor_category_equivalence.inverse.map f := rfl + variables (H : Group.{u}) instance [right_rigid_category V] : right_rigid_category (single_obj (H : Mon.{u}) ⥤ V) := diff --git a/src/representation_theory/Rep.lean b/src/representation_theory/Rep.lean index 376a1fc0447e1..f9d6375170a47 100644 --- a/src/representation_theory/Rep.lean +++ b/src/representation_theory/Rep.lean @@ -9,6 +9,7 @@ import algebra.category.Module.abelian import algebra.category.Module.colimits import algebra.category.Module.monoidal import algebra.category.Module.adjunctions +import category_theory.closed.functor_category /-! # `Rep k G` is the category of `k`-linear representations of `G`. @@ -39,8 +40,9 @@ 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} [comm_ring k] [monoid G] +variables {k G : Type u} [comm_ring k] +section +variables [monoid G] instance : has_coe_to_sort (Rep k G) (Type u) := concrete_category.has_coe_to_sort _ @@ -116,6 +118,97 @@ noncomputable def linearization_of_mul_action_iso (n : ℕ) : ≅ of_mul_action k G (fin n → G) := iso.refl _ end linearization +end +section +open Action +variables [group G] (A B C : Rep k G) + +noncomputable instance : monoidal_closed (Rep k G) := +monoidal_closed.of_equiv (functor_category_monoidal_equivalence _ _) + +/-- Explicit description of the 'internal Hom' `iHom(A, B)` of two representations `A, B`: +this is `F⁻¹(iHom(F(A), F(B)))`, where `F` is an equivalence +`Rep k G ≌ (single_obj G ⥤ Module k)`. Just used to prove `Rep.ihom_obj_ρ`. -/ +lemma ihom_obj_ρ_def : + ((ihom A).obj B).ρ = (functor_category_equivalence.inverse.obj + ((functor_category_equivalence.functor.obj A).closed_ihom.obj + (functor_category_equivalence.functor.obj B))).ρ := rfl + +/-- Given `k`-linear `G`-representations `(A, ρ₁), (B, ρ₂)`, the 'internal Hom' is the +representation on `Homₖ(A, B)` sending `g : G` and `f : A →ₗ[k] B` to `(ρ₂ g) ∘ₗ f ∘ₗ (ρ₁ g⁻¹)`. -/ +@[simp] lemma ihom_obj_ρ : + ((ihom A).obj B).ρ = A.ρ.lin_hom B.ρ := +begin + refine monoid_hom.ext (λ g, _), + simpa only [ihom_obj_ρ_def, functor_category_equivalence.inverse_obj_ρ_apply, + functor.closed_ihom_obj_map, ←functor.map_inv, single_obj.inv_as_inv], +end + +@[simp] lemma ihom_map_hom {B C : Rep k G} (f : B ⟶ C) : + ((ihom A).map f).hom = linear_map.llcomp k A B C f.hom := +rfl + +/-- Unfolds the unit in the adjunction `A ⊗ - ⊣ iHom(A, -)`; just used to prove +`Rep.ihom_coev_app_hom`. -/ +lemma ihom_coev_app_def : + (ihom.coev A).app B = functor_category_equivalence.unit_iso.hom.app B ≫ + functor_category_equivalence.inverse.map + ((functor_category_equivalence.functor.obj A).closed_unit.app _ ≫ + (functor_category_equivalence.functor.obj A).closed_ihom.map + ((functor_category_monoidal_equivalence (Module.{u} k) (Mon.of G)).μ A B)) := +rfl + +/-- Describes the unit in the adjunction `A ⊗ - ⊣ iHom(A, -)`; given another `k`-linear +`G`-representation `B,` the `k`-linear map underlying the resulting map `B ⟶ iHom(A, A ⊗ B)` is +given by flipping the arguments in the natural `k`-bilinear map `A →ₗ[k] B →ₗ[k] A ⊗ B`. -/ +@[simp] lemma ihom_coev_app_hom : + Action.hom.hom ((ihom.coev A).app B) = + (tensor_product.mk _ _ _).flip := +begin + refine linear_map.ext (λ x, linear_map.ext (λ y, _)), + simpa only [ihom_coev_app_def, functor.map_comp, comp_hom, + functor_category_equivalence.inverse_map_hom, functor.closed_ihom_map_app, + functor_category_monoidal_equivalence.μ_app], +end + +variables {A B C} + +/-- Given a `k`-linear `G`-representation `A`, the adjunction `A ⊗ - ⊣ iHom(A, -)` defines a +bijection `Hom(A ⊗ B, C) ≃ Hom(B, iHom(A, C))` for all `B, C`. Given `f : A ⊗ B ⟶ C`, this lemma +describes the `k`-linear map underlying `f`'s image under the bijection. It is given by currying the +`k`-linear map underlying `f`, giving a map `A →ₗ[k] B →ₗ[k] C`, then flipping the arguments. -/ +@[simp] lemma monoidal_closed_curry_hom (f : A ⊗ B ⟶ C) : + (monoidal_closed.curry f).hom = (tensor_product.curry f.hom).flip := +begin + rw [monoidal_closed.curry_eq, comp_hom, ihom_coev_app_hom], + refl, +end + +/-- Given a `k`-linear `G`-representation `A`, the adjunction `A ⊗ - ⊣ iHom(A, -)` defines a +bijection `Hom(A ⊗ B, C) ≃ Hom(B, iHom(A, C))` for all `B, C`. Given `f : B ⟶ iHom(A, C)`, this +lemma describes the `k`-linear map underlying `f`'s image under the bijection. It is given by +flipping the arguments of the `k`-linear map underlying `f`, giving a map `A →ₗ[k] B →ₗ[k] C`, then +uncurrying. -/ +@[simp] lemma monoidal_closed_uncurry_hom (f : B ⟶ (ihom A).obj C) : + (monoidal_closed.uncurry f).hom = tensor_product.uncurry _ _ _ _ f.hom.flip := +begin + simp only [monoidal_closed.of_equiv_uncurry_def, comp_inv_iso_inv_app, + monoidal_functor.comm_tensor_left_inv_app, comp_hom, + functor_category_monoidal_equivalence.inverse_map, functor_category_equivalence.inverse_map_hom, + functor_category_monoidal_equivalence.μ_iso_inv_app], + ext, + refl, +end + +/-- Describes the counit in the adjunction `A ⊗ - ⊣ iHom(A, -)`; given another `k`-linear +`G`-representation `B,` the `k`-linear map underlying the resulting morphism `A ⊗ iHom(A, B) ⟶ B` +is given by uncurrying the map `A →ₗ[k] (A →ₗ[k] B) →ₗ[k] B` defined by flipping the arguments in +the identity map on `Homₖ(A, B).` -/ +@[simp] lemma ihom_ev_app_hom : Action.hom.hom ((ihom.ev A).app B) = + tensor_product.uncurry _ _ _ _ linear_map.id.flip := +monoidal_closed_uncurry_hom _ + +end end Rep /-! From 72c366d0475675f1309d3027d3d7d47ee4423951 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 23 Mar 2023 23:41:06 +0000 Subject: [PATCH 0678/1291] feat(ring_theory/mv_polynomial/ideal): lemmas about monomial ideals (#18633) Inspired by [this Zulip message](https://leanprover.zulipchat.com/#narrow/stream/113489-new-members/topic/.F0.9D.94.BD.E2.82.82.5B.CE.B1.2C.20.CE.B2.2C.20.CE.B3.5D.20.2F.20.28.CE.B1.C2.B2.2C.20.CE.B2.C2.B2.2C.20.CE.B3.C2.B2.29/near/292327061). Instead of defining `ideal_ge'` directly, we show results about the more natural spelling with `ideal.span`. This also adds a handful of results about `dvd` on `mv_polynomial`. --- src/algebra/monoid_algebra/division.lean | 11 ++++ src/algebra/monoid_algebra/ideal.lean | 65 ++++++++++++++++++++++++ src/data/mv_polynomial/division.lean | 53 +++++++++++++++++++ src/ring_theory/mv_polynomial/ideal.lean | 58 +++++++++++++++++++++ 4 files changed, 187 insertions(+) create mode 100644 src/algebra/monoid_algebra/ideal.lean create mode 100644 src/ring_theory/mv_polynomial/ideal.lean diff --git a/src/algebra/monoid_algebra/division.lean b/src/algebra/monoid_algebra/division.lean index dcbd98f94a9f8..3e34e02a80b31 100644 --- a/src/algebra/monoid_algebra/division.lean +++ b/src/algebra/monoid_algebra/division.lean @@ -175,6 +175,17 @@ lemma mod_of_add_div_of (x : add_monoid_algebra k G) (g : G) : x %ᵒᶠ g + of' k G g * (x /ᵒᶠ g) = x := by rw [add_comm, div_of_add_mod_of] +lemma of'_dvd_iff_mod_of_eq_zero {x : add_monoid_algebra k G} {g : G} : + of' k G g ∣ x ↔ x %ᵒᶠ g = 0 := +begin + split, + { rintro ⟨x, rfl⟩, + rw of'_mul_mod_of }, + { intro h, + rw [←div_of_add_mod_of x g, h, add_zero], + exact dvd_mul_right _ _ }, +end + end end add_monoid_algebra diff --git a/src/algebra/monoid_algebra/ideal.lean b/src/algebra/monoid_algebra/ideal.lean new file mode 100644 index 0000000000000..046f3dff58f7f --- /dev/null +++ b/src/algebra/monoid_algebra/ideal.lean @@ -0,0 +1,65 @@ +/- +Copyright (c) 2023 Eric Wieser. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Eric Wieser +-/ + +import algebra.monoid_algebra.division +import ring_theory.ideal.basic + +/-! +# Lemmas about ideals of `monoid_algebra` and `add_monoid_algebra` +-/ + +variables {k A G : Type*} + +/-- If `x` belongs to the ideal generated by generators in `s`, then every element of the support of +`x` factors through an element of `s`. + +We could spell `∃ d, m = d * m` as `mul_opposite.op m' ∣ mul_opposite.op m` but this would be worse. +-/ +lemma monoid_algebra.mem_ideal_span_of_image + [monoid G] [semiring k] {s : set G} {x : monoid_algebra k G} : + x ∈ ideal.span (monoid_algebra.of k G '' s) ↔ ∀ m ∈ x.support, ∃ m' ∈ s, ∃ d, m = d * m' := +begin + let RHS : ideal (monoid_algebra k G) := + { carrier := {p | ∀ (m : G), m ∈ p.support → ∃ m' ∈ s, ∃ d, m = d * m'}, + add_mem' := λ x y hx hy m hm, by classical; + exact (finset.mem_union.1 $ finsupp.support_add hm).elim (hx m) (hy m), + zero_mem' := λ m hm, by cases hm, + smul_mem' := λ x y hy m hm, begin + replace hm := finset.mem_bUnion.mp (finsupp.support_sum hm), + obtain ⟨xm, hxm, hm⟩ := hm, + replace hm := finset.mem_bUnion.mp (finsupp.support_sum hm), + obtain ⟨ym, hym, hm⟩ := hm, + replace hm := finset.mem_singleton.mp (finsupp.support_single_subset hm), + obtain rfl := hm, + refine (hy _ hym).imp (λ sm, Exists.imp $ λ hsm, _), + rintros ⟨d, rfl⟩, + exact ⟨xm * d, (mul_assoc _ _ _).symm⟩, + end }, + change _ ↔ x ∈ RHS, + split, + { revert x, + refine ideal.span_le.2 _, + rintro _ ⟨i, hi, rfl⟩ m hm, + refine ⟨_, hi, 1, _⟩, + obtain rfl := finset.mem_singleton.mp (finsupp.support_single_subset hm), + exact (one_mul _).symm }, + { intros hx, + rw ←finsupp.sum_single x, + apply ideal.sum_mem _ (λ i hi, _), + obtain ⟨d, hd, d2, rfl⟩ := hx _ hi, + convert ideal.mul_mem_left _ (id $ finsupp.single d2 $ (x (d2 * d)) : monoid_algebra k G) _, + swap 3, + refine ideal.subset_span ⟨_, hd, rfl⟩, + rw [id.def, monoid_algebra.of_apply, monoid_algebra.single_mul_single, mul_one] }, +end + +/-- If `x` belongs to the ideal generated by generators in `s`, then every element of the support of +`x` factors additively through an element of `s`. +-/ +lemma add_monoid_algebra.mem_ideal_span_of'_image + [add_monoid A] [semiring k] {s : set A} {x : add_monoid_algebra k A} : + x ∈ ideal.span (add_monoid_algebra.of' k A '' s) ↔ ∀ m ∈ x.support, ∃ m' ∈ s, ∃ d, m = d + m' := +@monoid_algebra.mem_ideal_span_of_image k (multiplicative A) _ _ _ _ diff --git a/src/data/mv_polynomial/division.lean b/src/data/mv_polynomial/division.lean index 3dfa2febc6cba..9e44aca4ab738 100644 --- a/src/data/mv_polynomial/division.lean +++ b/src/data/mv_polynomial/division.lean @@ -110,8 +110,13 @@ lemma mod_monomial_add_div_monomial (x : mv_polynomial σ R) (s : σ →₀ ℕ) x %ᵐᵒⁿᵒᵐⁱᵃˡ s + monomial s 1 * (x /ᵐᵒⁿᵒᵐⁱᵃˡ s) = x := add_monoid_algebra.mod_of_add_div_of x s +lemma monomial_one_dvd_iff_mod_monomial_eq_zero {i : σ →₀ ℕ} {x : mv_polynomial σ R} : + monomial i (1 : R) ∣ x ↔ x %ᵐᵒⁿᵒᵐⁱᵃˡ i = 0 := +add_monoid_algebra.of'_dvd_iff_mod_of_eq_zero + end copied_declarations + section X_lemmas local infix ` /ᵐᵒⁿᵒᵐⁱᵃˡ `:70 := div_monomial @@ -149,6 +154,54 @@ lemma mod_monomial_add_div_monomial_single (x : mv_polynomial σ R) (i : σ) : x %ᵐᵒⁿᵒᵐⁱᵃˡ finsupp.single i 1 + X i * (x /ᵐᵒⁿᵒᵐⁱᵃˡ finsupp.single i 1) = x := mod_monomial_add_div_monomial _ _ +lemma X_dvd_iff_mod_monomial_eq_zero {i : σ} {x : mv_polynomial σ R} : + X i ∣ x ↔ x %ᵐᵒⁿᵒᵐⁱᵃˡ finsupp.single i 1 = 0 := +monomial_one_dvd_iff_mod_monomial_eq_zero + end X_lemmas +/-! ### Some results about dvd (`∣`) on `monomial` and `X` -/ + +lemma monomial_dvd_monomial {r s : R} {i j : σ →₀ ℕ} : + monomial i r ∣ monomial j s ↔ (s = 0 ∨ i ≤ j) ∧ r ∣ s := +begin + split, + { rintro ⟨x, hx⟩, + rw mv_polynomial.ext_iff at hx, + have hj := hx j, + have hi := hx i, + classical, + simp_rw [coeff_monomial, if_pos] at hj hi, + simp_rw [coeff_monomial_mul', if_pos] at hi hj, + split_ifs at hi hj with hi hi, + { exact ⟨or.inr hi, _, hj⟩ }, + { exact ⟨or.inl hj, hj.symm ▸ dvd_zero _⟩ } }, + { rintro ⟨h | hij, d, rfl⟩, + { simp_rw [h, monomial_zero, dvd_zero] }, + { refine ⟨monomial (j - i) d, _⟩, + rw [monomial_mul, add_tsub_cancel_of_le hij] } } +end + +@[simp] lemma monomial_one_dvd_monomial_one [nontrivial R] {i j : σ →₀ ℕ} : + monomial i (1 : R) ∣ monomial j 1 ↔ i ≤ j := +begin + rw monomial_dvd_monomial, + simp_rw [one_ne_zero, false_or, dvd_rfl, and_true], +end + +@[simp] lemma X_dvd_X [nontrivial R] {i j : σ} : + (X i : mv_polynomial σ R) ∣ (X j : mv_polynomial σ R) ↔ i = j := +begin + refine monomial_one_dvd_monomial_one.trans _, + simp_rw [finsupp.single_le_iff, nat.one_le_iff_ne_zero, finsupp.single_apply_ne_zero, + ne.def, one_ne_zero, not_false_iff, and_true], +end + +@[simp] lemma X_dvd_monomial {i : σ} {j : σ →₀ ℕ} {r : R} : + (X i : mv_polynomial σ R) ∣ monomial j r ↔ (r = 0 ∨ j i ≠ 0) := +begin + refine monomial_dvd_monomial.trans _, + simp_rw [one_dvd, and_true, finsupp.single_le_iff, nat.one_le_iff_ne_zero], +end + end mv_polynomial diff --git a/src/ring_theory/mv_polynomial/ideal.lean b/src/ring_theory/mv_polynomial/ideal.lean new file mode 100644 index 0000000000000..2faf6d7280dcf --- /dev/null +++ b/src/ring_theory/mv_polynomial/ideal.lean @@ -0,0 +1,58 @@ +/- +Copyright (c) 2023 Eric Wieser. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Eric Wieser +-/ +import algebra.monoid_algebra.ideal +import data.mv_polynomial.division + +/-! +# Lemmas about ideals of `mv_polynomial` + +Notably this contains results about monomial ideals. + +## Main results + +* `mv_polynomial.mem_ideal_span_monomial_image` +* `mv_polynomial.mem_ideal_span_X_image` +-/ + +variables {σ R : Type*} + +namespace mv_polynomial +variables [comm_semiring R] + + +/-- `x` is in a monomial ideal generated by `s` iff every element of of its support dominates one of +the generators. Note that `si ≤ xi` is analogous to saying that the monomial corresponding to `si` +divides the monomial corresponding to `xi`. -/ +lemma mem_ideal_span_monomial_image + {x : mv_polynomial σ R} {s : set (σ →₀ ℕ)} : + x ∈ ideal.span ((λ s, monomial s (1 : R)) '' s) ↔ ∀ xi ∈ x.support, ∃ si ∈ s, si ≤ xi := +begin + refine add_monoid_algebra.mem_ideal_span_of'_image.trans _, + simp_rw [le_iff_exists_add, add_comm], + refl, +end + +lemma mem_ideal_span_monomial_image_iff_dvd {x : mv_polynomial σ R} {s : set (σ →₀ ℕ)} : + x ∈ ideal.span ((λ s, monomial s (1 : R)) '' s) ↔ + ∀ xi ∈ x.support, ∃ si ∈ s, monomial si 1 ∣ monomial xi (x.coeff xi) := +begin + refine mem_ideal_span_monomial_image.trans (forall₂_congr $ λ xi hxi, _), + simp_rw [monomial_dvd_monomial, one_dvd, and_true, mem_support_iff.mp hxi, false_or], +end + +/-- `x` is in a monomial ideal generated by variables `X` iff every element of of its support +has a component in `s`. -/ +lemma mem_ideal_span_X_image {x : mv_polynomial σ R} {s : set σ} : + x ∈ ideal.span (mv_polynomial.X '' s : set (mv_polynomial σ R)) ↔ + ∀ m ∈ x.support, ∃ i ∈ s, (m : σ →₀ ℕ) i ≠ 0 := +begin + have := @mem_ideal_span_monomial_image σ R _ _ ((λ i, finsupp.single i 1) '' s), + rw set.image_image at this, + refine this.trans _, + simp [nat.one_le_iff_ne_zero], +end + +end mv_polynomial From e9ce88cd0d54891c714c604076084f763dd480ed Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 24 Mar 2023 01:43:32 +0000 Subject: [PATCH 0679/1291] feat(order/upper_lower/basic): The upper closure is bounded below (#18637) and its lower bounds are the same as those of the original set. --- src/order/upper_lower/basic.lean | 19 +++++++++++++++++++ 1 file changed, 19 insertions(+) diff --git a/src/order/upper_lower/basic.lean b/src/order/upper_lower/basic.lean index bfc501061dd55..d95fe66e45e02 100644 --- a/src/order/upper_lower/basic.lean +++ b/src/order/upper_lower/basic.lean @@ -839,6 +839,25 @@ begin exact (upper_set.upper _).ord_connected.inter (lower_set.lower _).ord_connected, end +@[simp] lemma upper_bounds_lower_closure : + upper_bounds (lower_closure s : set α) = upper_bounds s := +(upper_bounds_mono_set subset_lower_closure).antisymm $ λ a ha b ⟨c, hc, hcb⟩, hcb.trans $ ha hc + +@[simp] lemma lower_bounds_upper_closure : + lower_bounds (upper_closure s : set α) = lower_bounds s := +(lower_bounds_mono_set subset_upper_closure).antisymm $ λ a ha b ⟨c, hc, hcb⟩, (ha hc).trans hcb + +@[simp] lemma bdd_above_lower_closure : bdd_above (lower_closure s : set α) ↔ bdd_above s := +by simp_rw [bdd_above, upper_bounds_lower_closure] + +@[simp] lemma bdd_below_upper_closure : bdd_below (upper_closure s : set α) ↔ bdd_below s := +by simp_rw [bdd_below, lower_bounds_upper_closure] + +alias bdd_above_lower_closure ↔ bdd_above.of_lower_closure bdd_above.lower_closure +alias bdd_below_upper_closure ↔ bdd_below.of_upper_closure bdd_below.upper_closure + +attribute [protected] bdd_above.lower_closure bdd_below.upper_closure + end closure /-! ### Product -/ From 97eab48559068f3d6313da387714ef25768fb730 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Fri, 24 Mar 2023 07:37:25 +0000 Subject: [PATCH 0680/1291] chore(*): add mathlib4 synchronization comments (#18642) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.char_p.algebra` * `analysis.normed_space.M_structure` * `category_theory.groupoid.vertex_group` * `category_theory.limits.shapes.multiequalizer` * `category_theory.limits.shapes.reflexive` * `combinatorics.derangements.basic` * `combinatorics.simple_graph.degree_sum` * `data.polynomial.denoms_clearable` * `data.polynomial.ring_division` * `data.set.sups` * `linear_algebra.affine_space.affine_equiv` * `number_theory.pythagorean_triples` * `number_theory.zsqrtd.basic` * `number_theory.zsqrtd.to_real` * `ring_theory.localization.num_denom` * `topology.path_connected` --- src/algebra/char_p/algebra.lean | 3 +++ src/analysis/normed_space/M_structure.lean | 3 +++ src/category_theory/groupoid/vertex_group.lean | 3 +++ src/category_theory/limits/shapes/multiequalizer.lean | 3 +++ src/category_theory/limits/shapes/reflexive.lean | 3 +++ src/combinatorics/derangements/basic.lean | 3 +++ src/combinatorics/simple_graph/degree_sum.lean | 3 +++ src/data/polynomial/denoms_clearable.lean | 3 +++ src/data/polynomial/ring_division.lean | 3 +++ src/data/set/sups.lean | 3 +++ src/linear_algebra/affine_space/affine_equiv.lean | 3 +++ src/number_theory/pythagorean_triples.lean | 3 +++ src/number_theory/zsqrtd/basic.lean | 3 +++ src/number_theory/zsqrtd/to_real.lean | 3 +++ src/ring_theory/localization/num_denom.lean | 3 +++ src/topology/path_connected.lean | 3 +++ 16 files changed, 48 insertions(+) diff --git a/src/algebra/char_p/algebra.lean b/src/algebra/char_p/algebra.lean index 0748eeb869bcb..2273ed60cc6ee 100644 --- a/src/algebra/char_p/algebra.lean +++ b/src/algebra/char_p/algebra.lean @@ -11,6 +11,9 @@ import algebra.free_algebra /-! # Characteristics of algebras +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we describe the characteristic of `R`-algebras. In particular we are interested in the characteristic of free algebras over `R` diff --git a/src/analysis/normed_space/M_structure.lean b/src/analysis/normed_space/M_structure.lean index 6f38dbc4af335..fceb3e6b42f1c 100644 --- a/src/analysis/normed_space/M_structure.lean +++ b/src/analysis/normed_space/M_structure.lean @@ -10,6 +10,9 @@ import analysis.normed.group.basic /-! # M-structure +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A projection P on a normed space X is said to be an L-projection (`is_Lprojection`) if, for all `x` in `X`, $\|x\| = \|P x\| + \|(1 - P) x\|$. diff --git a/src/category_theory/groupoid/vertex_group.lean b/src/category_theory/groupoid/vertex_group.lean index 58c8c6e968a1b..d4bbadc7c5615 100644 --- a/src/category_theory/groupoid/vertex_group.lean +++ b/src/category_theory/groupoid/vertex_group.lean @@ -13,6 +13,9 @@ import combinatorics.quiver.path /-! # Vertex group +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the vertex group (*aka* isotropy group) of a groupoid at a vertex. ## Implementation notes diff --git a/src/category_theory/limits/shapes/multiequalizer.lean b/src/category_theory/limits/shapes/multiequalizer.lean index 094b360ab504d..aa905aa8b2ca0 100644 --- a/src/category_theory/limits/shapes/multiequalizer.lean +++ b/src/category_theory/limits/shapes/multiequalizer.lean @@ -11,6 +11,9 @@ import category_theory.limits.cone_category # Multi-(co)equalizers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A *multiequalizer* is an equalizer of two morphisms between two products. Since both products and equalizers are limits, such an object is again a limit. This file provides the diagram whose limit is indeed such an object. diff --git a/src/category_theory/limits/shapes/reflexive.lean b/src/category_theory/limits/shapes/reflexive.lean index 9415a43050ad7..96a14fb7045ad 100644 --- a/src/category_theory/limits/shapes/reflexive.lean +++ b/src/category_theory/limits/shapes/reflexive.lean @@ -9,6 +9,9 @@ import category_theory.limits.shapes.kernel_pair /-! # Reflexive coequalizers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define reflexive pairs as a pair of morphisms which have a common section. We say a category has reflexive coequalizers if it has coequalizers of all reflexive pairs. Reflexive coequalizers often enjoy nicer properties than general coequalizers, and feature heavily diff --git a/src/combinatorics/derangements/basic.lean b/src/combinatorics/derangements/basic.lean index 6d0cce36db9ec..efac282381219 100644 --- a/src/combinatorics/derangements/basic.lean +++ b/src/combinatorics/derangements/basic.lean @@ -11,6 +11,9 @@ import logic.equiv.option /-! # Derangements on types +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define `derangements α`, the set of derangements on a type `α`. We also define some equivalences involving various subtypes of `perm α` and `derangements α`: diff --git a/src/combinatorics/simple_graph/degree_sum.lean b/src/combinatorics/simple_graph/degree_sum.lean index f6ec91b9e4a21..06e9a8e6aa42e 100644 --- a/src/combinatorics/simple_graph/degree_sum.lean +++ b/src/combinatorics/simple_graph/degree_sum.lean @@ -11,6 +11,9 @@ import data.zmod.parity /-! # Degree-sum formula and handshaking lemma +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The degree-sum formula is that the sum of the degrees of the vertices in a finite graph is equal to twice the number of edges. The handshaking lemma, a corollary, is that the number of odd-degree vertices is even. diff --git a/src/data/polynomial/denoms_clearable.lean b/src/data/polynomial/denoms_clearable.lean index bf5a0d881afeb..39db6d8a3f533 100644 --- a/src/data/polynomial/denoms_clearable.lean +++ b/src/data/polynomial/denoms_clearable.lean @@ -9,6 +9,9 @@ import data.polynomial.eval /-! # Denominators of evaluation of polynomials at ratios +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Let `i : R → K` be a homomorphism of semirings. Assume that `K` is commutative. If `a` and `b` are elements of `R` such that `i b ∈ K` is invertible, then for any polynomial `f ∈ R[X]` the "mathematical" expression `b ^ f.nat_degree * f (a / b) ∈ K` is in diff --git a/src/data/polynomial/ring_division.lean b/src/data/polynomial/ring_division.lean index c5e48fc31eb63..f64a1c21d0d01 100644 --- a/src/data/polynomial/ring_division.lean +++ b/src/data/polynomial/ring_division.lean @@ -13,6 +13,9 @@ import algebra.polynomial.big_operators /-! # Theory of univariate polynomials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file starts looking like the ring theory of $ R[X] $ ## Main definitions diff --git a/src/data/set/sups.lean b/src/data/set/sups.lean index d06e854a18b3e..a03506695192e 100644 --- a/src/data/set/sups.lean +++ b/src/data/set/sups.lean @@ -9,6 +9,9 @@ import order.upper_lower.basic /-! # Set family operations +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines a few binary operations on `set α` for use in set family combinatorics. ## Main declarations diff --git a/src/linear_algebra/affine_space/affine_equiv.lean b/src/linear_algebra/affine_space/affine_equiv.lean index 766c01c3a433b..c81fc970ecac9 100644 --- a/src/linear_algebra/affine_space/affine_equiv.lean +++ b/src/linear_algebra/affine_space/affine_equiv.lean @@ -10,6 +10,9 @@ import algebra.invertible /-! # Affine equivalences +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define `affine_equiv k P₁ P₂` (notation: `P₁ ≃ᵃ[k] P₂`) to be the type of affine equivalences between `P₁` and `P₂, i.e., equivalences such that both forward and inverse maps are affine maps. diff --git a/src/number_theory/pythagorean_triples.lean b/src/number_theory/pythagorean_triples.lean index 089220078ac61..da59fbbe9b361 100644 --- a/src/number_theory/pythagorean_triples.lean +++ b/src/number_theory/pythagorean_triples.lean @@ -14,6 +14,9 @@ import data.zmod.basic /-! # Pythagorean Triples +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The main result is the classification of Pythagorean triples. The final result is for general Pythagorean triples. It follows from the more interesting relatively prime case. We use the "rational parametrization of the circle" method for the proof. The parametrization maps the point diff --git a/src/number_theory/zsqrtd/basic.lean b/src/number_theory/zsqrtd/basic.lean index ea3dbca7b990f..ded1c18c7e42a 100644 --- a/src/number_theory/zsqrtd/basic.lean +++ b/src/number_theory/zsqrtd/basic.lean @@ -10,6 +10,9 @@ import algebra.star.unitary /-! # ℤ[√d] +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The ring of integers adjoined with a square root of `d : ℤ`. After defining the norm, we show that it is a linearly ordered commutative ring, diff --git a/src/number_theory/zsqrtd/to_real.lean b/src/number_theory/zsqrtd/to_real.lean index aeb45b6ec98a8..8277868f99867 100644 --- a/src/number_theory/zsqrtd/to_real.lean +++ b/src/number_theory/zsqrtd/to_real.lean @@ -9,6 +9,9 @@ import number_theory.zsqrtd.basic /-! # Image of `zsqrtd` in `ℝ` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines `zsqrtd.to_real` and related lemmas. It is in a separate file to avoid pulling in all of `data.real` into `data.zsqrtd`. -/ diff --git a/src/ring_theory/localization/num_denom.lean b/src/ring_theory/localization/num_denom.lean index 2a33e89e230b2..1b1ff85f10758 100644 --- a/src/ring_theory/localization/num_denom.lean +++ b/src/ring_theory/localization/num_denom.lean @@ -10,6 +10,9 @@ import ring_theory.unique_factorization_domain /-! # Numerator and denominator in a localization +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Implementation notes See `src/ring_theory/localization/basic.lean` for a design overview. diff --git a/src/topology/path_connected.lean b/src/topology/path_connected.lean index 76e5e9542fd13..3e52c69b1f571 100644 --- a/src/topology/path_connected.lean +++ b/src/topology/path_connected.lean @@ -11,6 +11,9 @@ import topology.unit_interval /-! # Path connectedness +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions In the file the unit interval `[0, 1]` in `ℝ` is denoted by `I`, and `X` is a topological space. From 540b766a64a8cc1e4b013f43a31e3b0b09787937 Mon Sep 17 00:00:00 2001 From: Rob Lewis Date: Fri, 24 Mar 2023 12:53:51 +0000 Subject: [PATCH 0681/1291] feat(tactic/linear_combination): allow `linear_combination with { exponent := n }` (#15428) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit #15425 will solve problems where the goal is not a linear combination of hypotheses, but a *power* of the goal is. This PR provides a more natural certificate for these proofs. Writing `linear_combination ... with { exponent := 2 }` will square the goal before subtracting the linear combination. In principle this could be useful even outside of `polyrith`. Co-authored-by: Rob Lewis Co-authored-by: Yaël Dillies --- src/tactic/linear_combination.lean | 19 +++++++++++++++++++ test/linear_combination.lean | 25 +++++++++++++++++++++++++ 2 files changed, 44 insertions(+) diff --git a/src/tactic/linear_combination.lean b/src/tactic/linear_combination.lean index 306bfd6ea7439..48173a2f005a9 100644 --- a/src/tactic/linear_combination.lean +++ b/src/tactic/linear_combination.lean @@ -73,6 +73,7 @@ checking if the weighted sum is equivalent to the goal (when `normalize` is `tt` meta structure linear_combination_config : Type := (normalize : bool := tt) (normalization_tactic : tactic unit := `[ring_nf SOP]) +(exponent : ℕ := 1) /-! ### Part 1: Multiplying Equations by Constants and Adding Them Together -/ @@ -266,6 +267,17 @@ This tactic only should be used when the target's type is an equality whose meta def set_goal_to_hleft_sub_tleft (hsum_on_left : expr) : tactic unit := do to_expr ``(eq_zero_of_sub_eq_zero %%hsum_on_left) >>= apply, skip +/-- +If an exponent `n` is provided, changes the goal from `t = 0` to `t^n = 0`. +* Input: + * `exponent : ℕ`, the power to raise the goal by. If `1`, this tactic is a no-op. + +* Output: N/A +-/ +meta def raise_goal_to_power : ℕ → tactic unit +| 1 := skip +| n := refine ``(@pow_eq_zero _ _ _ _ %%`(n) _) + /-- This tactic attempts to prove the goal by normalizing the target if the `normalize` field of the given configuration is true. @@ -314,6 +326,7 @@ do hsum ← make_sum_of_hyps ext h_eqs coeffs, hsum_on_left ← move_to_left_side hsum, move_target_to_left_side, + raise_goal_to_power config.exponent, set_goal_to_hleft_sub_tleft hsum_on_left, normalize_if_desired config @@ -354,6 +367,9 @@ setup_tactic_parser configuration is set to false, then the tactic will simply set the user up to prove their target using the linear combination instead of normalizing the subtraction. +Users may provide an optional `with { exponent := n }`. This will raise the goal to the power `n` + before subtracting the linear combination. + Note: The left and right sides of all the equalities should have the same type, and the coefficients should also have this type. There must be instances of `has_mul` and `add_group` for this type. @@ -406,6 +422,9 @@ begin norm_cast end +example (x y z : ℚ) (h : x = y) (h2 : x * y = 0) : x + y*z = 0 := +by linear_combination (-y * z ^ 2 + x) * h + (z ^ 2 + 2 * z + 1) * h2 with { exponent := 2 } + constants (qc : ℚ) (hqc : qc = 2*qc) example (a b : ℚ) (h : ∀ p q : ℚ, p = q) : 3*a + qc = 3*b + 2*qc := diff --git a/test/linear_combination.lean b/test/linear_combination.lean index e318dc1a394ea..c1f0921b46dee 100644 --- a/test/linear_combination.lean +++ b/test/linear_combination.lean @@ -199,6 +199,31 @@ end example {x y z w : ℤ} (h₁ : 3 * x = 4 + y) (h₂ : x + 2 * y = 1) : z + w = w + z := by linear_combination with {normalization_tactic := `[simp [add_comm]]} +/-! ### Cases with exponent -/ + +example (x y z : ℚ) (h : x = y) (h2 : x * y = 0) : x + y*z = 0 := +by linear_combination (-y * z ^ 2 + x) * h + (z ^ 2 + 2 * z + 1) * h2 with {exponent := 2} + +example (x y z : ℚ) (h : x = y) (h2 : x * y = 0) : y*z = -x := +begin + linear_combination (-y * z ^ 2 + x) * h + (z ^ 2 + 2 * z + 1) * h2 + with {exponent := 2, normalize := ff}, + ring +end + +example (K : Type) + [field K] + [char_zero K] + {x y z : K} + (h₂ : y ^ 3 + x * (3 * z ^ 2) = 0) + (h₁ : x ^ 3 + z * (3 * y ^ 2) = 0) + (h₀ : y * (3 * x ^ 2) + z ^ 3 = 0) + (h : x ^ 3 * y + y ^ 3 * z + z ^ 3 * x = 0) : + x = 0 := +by linear_combination 2 * y * z ^ 2 * h₂ / 7 + (x ^ 3 - y ^ 2 * z / 7) * h₁ - + x * y * z * h₀ + y * z * h / 7 with {exponent := 6} + + /-! ### Cases where the goal is not closed -/ example (x y : ℚ) (h1 : x + y = 3) (h2 : 3*x = 7) : From 945bc74ecd6c7435f33e080af142c1cfe8d2e289 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 24 Mar 2023 14:33:22 +0000 Subject: [PATCH 0682/1291] chore(data/matrix/kronecker): make the `R` argument implicit (#18624) This was copied erroneously from the `tensor_product` section, where an explicit `R` _is_ needed. --- src/data/matrix/kronecker.lean | 2 -- 1 file changed, 2 deletions(-) diff --git a/src/data/matrix/kronecker.lean b/src/data/matrix/kronecker.lean index 5dc1f8bfeada9..dfe8fa14ac296 100644 --- a/src/data/matrix/kronecker.lean +++ b/src/data/matrix/kronecker.lean @@ -246,8 +246,6 @@ end kronecker_map section kronecker -variables (R) - open_locale matrix /-- The Kronecker product. This is just a shorthand for `kronecker_map (*)`. Prefer the notation From d11893b411025250c8e61ff2f12ccbd7ee35ab15 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 24 Mar 2023 18:27:09 +0000 Subject: [PATCH 0683/1291] fix(*): missing universe polymorphism (#18644) These are all just typo fixes, no proof adaptations. This deliberately leaves alone things related to category theory and algebraic geometry, as there the lack of polymorphism is likely deliberate. --- src/analysis/normed_space/M_structure.lean | 2 +- src/data/list/func.lean | 4 ++-- src/field_theory/separable_degree.lean | 4 ++-- src/linear_algebra/bilinear_form.lean | 2 +- src/logic/equiv/list.lean | 2 +- src/measure_theory/integral/circle_transform.lean | 2 +- src/model_theory/elementary_maps.lean | 4 ++-- src/number_theory/legendre_symbol/gauss_sum.lean | 2 +- src/number_theory/legendre_symbol/quadratic_char.lean | 4 ++-- 9 files changed, 13 insertions(+), 13 deletions(-) diff --git a/src/analysis/normed_space/M_structure.lean b/src/analysis/normed_space/M_structure.lean index fceb3e6b42f1c..82bcc2458dad4 100644 --- a/src/analysis/normed_space/M_structure.lean +++ b/src/analysis/normed_space/M_structure.lean @@ -62,7 +62,7 @@ M-summand, M-projection, L-summand, L-projection, M-ideal, M-structure -/ variables (X : Type*) [normed_add_comm_group X] -variables {M : Type} [ring M] [module M X] +variables {M : Type*} [ring M] [module M X] /-- A projection on a normed space `X` is said to be an L-projection if, for all `x` in `X`, $\|x\| = \|P x\| + \|(1 - P) x\|$. diff --git a/src/data/list/func.lean b/src/data/list/func.lean index f075c0c84e28f..b43eda69369a4 100644 --- a/src/data/list/func.lean +++ b/src/data/list/func.lean @@ -354,7 +354,7 @@ by {apply get_pointwise, apply sub_zero} (sub xs ys).length = max xs.length ys.length := @length_pointwise α α α ⟨0⟩ ⟨0⟩ _ _ _ -@[simp] lemma nil_sub {α : Type} [add_group α] +@[simp] lemma nil_sub {α : Type*} [add_group α] (as : list α) : sub [] as = neg as := begin rw [sub, nil_pointwise], @@ -362,7 +362,7 @@ begin rw [zero_sub] end -@[simp] lemma sub_nil {α : Type} [add_group α] +@[simp] lemma sub_nil {α : Type*} [add_group α] (as : list α) : sub as [] = as := begin rw [sub, pointwise_nil], diff --git a/src/field_theory/separable_degree.lean b/src/field_theory/separable_degree.lean index 31645af5673f8..bc897af345ed3 100644 --- a/src/field_theory/separable_degree.lean +++ b/src/field_theory/separable_degree.lean @@ -41,7 +41,7 @@ open_locale classical polynomial section comm_semiring -variables {F : Type} [comm_semiring F] (q : ℕ) +variables {F : Type*} [comm_semiring F] (q : ℕ) /-- A separable contraction of a polynomial `f` is a separable polynomial `g` such that `g(x^(q^m)) = f(x)` for some `m : ℕ`.-/ @@ -86,7 +86,7 @@ end comm_semiring section field -variables {F : Type} [field F] +variables {F : Type*} [field F] variables (q : ℕ) {f : F[X]} (hf : has_separable_contraction q f) /-- Every irreducible polynomial can be contracted to a separable polynomial. diff --git a/src/linear_algebra/bilinear_form.lean b/src/linear_algebra/bilinear_form.lean index 1f4d13d0f31ab..a76ba207c864f 100644 --- a/src/linear_algebra/bilinear_form.lean +++ b/src/linear_algebra/bilinear_form.lean @@ -405,7 +405,7 @@ end equiv_lin namespace linear_map -variables {R' : Type} [comm_semiring R'] [algebra R' R] [module R' M] [is_scalar_tower R' R M] +variables {R' : Type*} [comm_semiring R'] [algebra R' R] [module R' M] [is_scalar_tower R' R M] /-- Apply a linear map on the output of a bilinear form. -/ @[simps] diff --git a/src/logic/equiv/list.lean b/src/logic/equiv/list.lean index d37847f6caffc..8a960fc413368 100644 --- a/src/logic/equiv/list.lean +++ b/src/logic/equiv/list.lean @@ -347,7 +347,7 @@ def list_unit_equiv : list unit ≃ ℕ := def list_nat_equiv_nat : list ℕ ≃ ℕ := denumerable.eqv _ /-- If `α` is equivalent to `ℕ`, then `list α` is equivalent to `α`. -/ -def list_equiv_self_of_equiv_nat {α : Type} (e : α ≃ ℕ) : list α ≃ α := +def list_equiv_self_of_equiv_nat {α : Type*} (e : α ≃ ℕ) : list α ≃ α := calc list α ≃ list ℕ : list_equiv_of_equiv e ... ≃ ℕ : list_nat_equiv_nat ... ≃ α : e.symm diff --git a/src/measure_theory/integral/circle_transform.lean b/src/measure_theory/integral/circle_transform.lean index 1356aa0a5464c..e721eca6870d0 100644 --- a/src/measure_theory/integral/circle_transform.lean +++ b/src/measure_theory/integral/circle_transform.lean @@ -22,7 +22,7 @@ open_locale interval real noncomputable theory -variables {E : Type} [normed_add_comm_group E] [normed_space ℂ E] (R : ℝ) (z w : ℂ) +variables {E : Type*} [normed_add_comm_group E] [normed_space ℂ E] (R : ℝ) (z w : ℂ) namespace complex diff --git a/src/model_theory/elementary_maps.lean b/src/model_theory/elementary_maps.lean index d62105d2b95d0..3db612f895343 100644 --- a/src/model_theory/elementary_maps.lean +++ b/src/model_theory/elementary_maps.lean @@ -59,7 +59,7 @@ instance fun_like : fun_like (M ↪ₑ[L] N) M (λ _, N) := instance : has_coe_to_fun (M ↪ₑ[L] N) (λ _, M → N) := fun_like.has_coe_to_fun -@[simp] lemma map_bounded_formula (f : M ↪ₑ[L] N) {α : Type} {n : ℕ} +@[simp] lemma map_bounded_formula (f : M ↪ₑ[L] N) {α : Type*} {n : ℕ} (φ : L.bounded_formula α n) (v : α → M) (xs : fin n → M) : φ.realize (f ∘ v) (f ∘ xs) ↔ φ.realize v xs := begin @@ -80,7 +80,7 @@ begin bounded_formula.realize_restrict_free_var set.subset.rfl], end -@[simp] lemma map_formula (f : M ↪ₑ[L] N) {α : Type} (φ : L.formula α) (x : α → M) : +@[simp] lemma map_formula (f : M ↪ₑ[L] N) {α : Type*} (φ : L.formula α) (x : α → M) : φ.realize (f ∘ x) ↔ φ.realize x := by rw [formula.realize, formula.realize, ← f.map_bounded_formula, unique.eq_default (f ∘ default)] diff --git a/src/number_theory/legendre_symbol/gauss_sum.lean b/src/number_theory/legendre_symbol/gauss_sum.lean index 5a1ae965e601a..17b8fa9862a53 100644 --- a/src/number_theory/legendre_symbol/gauss_sum.lean +++ b/src/number_theory/legendre_symbol/gauss_sum.lean @@ -207,7 +207,7 @@ end /-- When `F` and `F'` are finite fields and `χ : F → F'` is a nontrivial quadratic character, then `(χ(-1) * #F)^(#F'/2) = χ(#F')`. -/ -lemma char.card_pow_card {F : Type} [field F] [fintype F] {F' : Type} [field F'] [fintype F'] +lemma char.card_pow_card {F : Type*} [field F] [fintype F] {F' : Type*} [field F'] [fintype F'] {χ : mul_char F F'} (hχ₁ : is_nontrivial χ) (hχ₂ : is_quadratic χ) (hch₁ : ring_char F' ≠ ring_char F) (hch₂ : ring_char F' ≠ 2) : (χ (-1) * fintype.card F) ^ (fintype.card F' / 2) = χ (fintype.card F') := diff --git a/src/number_theory/legendre_symbol/quadratic_char.lean b/src/number_theory/legendre_symbol/quadratic_char.lean index 9e6d5d17c9fde..b49542517b746 100644 --- a/src/number_theory/legendre_symbol/quadratic_char.lean +++ b/src/number_theory/legendre_symbol/quadratic_char.lean @@ -287,7 +287,7 @@ section special_values open zmod mul_char -variables {F : Type} [field F] [fintype F] +variables {F : Type*} [field F] [fintype F] /-- The value of the quadratic character at `-1` -/ lemma quadratic_char_neg_one [decidable_eq F] (hF : ring_char F ≠ 2) : @@ -382,7 +382,7 @@ end /-- The relation between the values of the quadratic character of one field `F` at the cardinality of another field `F'` and of the quadratic character of `F'` at the cardinality of `F`. -/ -lemma quadratic_char_card_card [decidable_eq F] (hF : ring_char F ≠ 2) {F' : Type} [field F'] +lemma quadratic_char_card_card [decidable_eq F] (hF : ring_char F ≠ 2) {F' : Type*} [field F'] [fintype F'] [decidable_eq F'] (hF' : ring_char F' ≠ 2) (h : ring_char F' ≠ ring_char F) : quadratic_char F (fintype.card F') = quadratic_char F' (quadratic_char F (-1) * fintype.card F) := begin From 46b633fd842bef9469441c0209906f6dddd2b4f5 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 24 Mar 2023 21:41:31 +0000 Subject: [PATCH 0684/1291] refactor(analysis/inner_product_space/basic): do not extend normed_add_comm_group (#18583) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This replaces `[inner_product_space K V]` with `[normed_add_comm_group V] [inner_product_space K V]`. Before this PR, we had `inner_product_space K V extends normed_add_comm_group V`. At first glance this is a rather strange arrangement; we certainly don't have `module R V extends add_comm_group V`. The argument usually goes that `Derived A B extends Base B` is unsafe, because the `Derived.toBase` instance has no way of knowing which `A` to look for. For `inner_product_space K V` we argued that this problem doesn't apply, as `is_R_or_C K` means that in practice there are only two possible values of `K`. This is indeed enough to stop typeclass search grinding to a halt. The motivation for the old design is described by @dupuisf [on Zulip](https://leanprover.zulipchat.com/#narrow/stream/116395-maths/topic/Dangerous.20instance.20vs.20elaboration.20problems/near/210594971): > However, when I do (the contents of this PR), I run into some very annoying elaboration issues where I have to constantly spoonfeed it `𝕜` and/or `α` in lemmas that rewrite norms in terms of inner products, even though the relevant instance is directly present in the context. While it may ease elaboration issues, there are a number of new problems that `inner_product_space K V extends normed_add_comm_group V` causes: 1. It is confusing when compared to all other typeclasses. After being told to write ```lean variables [add_comm_group U] [module R U] variables [normed_add_comm_group V] [normed_space R V] ``` it is natural to assume that the inner product space version would be ```lean variables [normed_add_comm_group W] [inner_product_space K W] ``` But writing this leads to `W` having two unrelated addition operations! 2. It doesn't allow a space `W` to be an inner product space over both R and C. For a (normed) module, you can write ```lean variables [add_comm_group U] [module R U] [module C U] variables [normed_add_comm_group V] [normed_space R V] [normed_space C V] ``` but writing `[inner_product_space R W] [inner_product_space C W]` again puts two unrelated `+` operators on `V` 3. It doesn't compose with other normed typeclasses. We can talk about a (normed) module with multiplication as ```lean variables [ring U] [module R U] variables [normed_ring V] [normed_space K V] ``` but once again, typing `[normed_ring W] [inner_product_space K W]` gives us two unrelated `+` operators. This prevents us generalizing over `real`, `complex`, and `quaternion`. We could work around this by defining variants of `inner_product_space` for `normed_ring A`, `normed_comm_ring A`, `normed_division_ring A`, `normed_field A`, but this doesn't scale well. If we do things "the module way" then we only need one new typeclass instead of four, ```lean class inner_product_algebra (K A) [is_R_or_C K] [normed_ring A] extends normed_algebra K A, inner_product_space K A ``` 5. The "`is_R_or_C` makes it OK" argument stops working if we generalize to [quaternionic inner product spaces](https://link.springer.com/chapter/10.1007/978-94-009-3713-0_13) The majority of this PR is adding `[normed_add_comm_group _]` to every `variables` line about `inner_product_space`. The rest is specifying the scalar manually where previously unification would find it for us. [This Zulip thread](https://leanprover.zulipchat.com/#narrow/stream/116395-maths/topic/inner_product_space.20and.20normed_algebra/near/341848831) has some initial discussion about this change. --- .../100-theorems-list/57_herons_formula.lean | 4 +- archive/imo/imo2019_q2.lean | 3 +- .../calculus/bump_function_inner.lean | 4 +- .../calculus/conformal/inner_product.lean | 4 +- src/analysis/convex/cone/basic.lean | 2 +- src/analysis/fourier/add_circle.lean | 2 +- src/analysis/inner_product_space/adjoint.lean | 9 +- src/analysis/inner_product_space/basic.lean | 152 +++++++++--------- .../inner_product_space/calculus.lean | 93 +++++------ .../conformal_linear_map.lean | 4 +- src/analysis/inner_product_space/dual.lean | 2 +- .../gram_schmidt_ortho.lean | 2 +- .../inner_product_space/l2_space.lean | 10 +- .../inner_product_space/lax_milgram.lean | 4 +- .../inner_product_space/orientation.lean | 5 +- src/analysis/inner_product_space/pi_L2.lean | 28 ++-- .../inner_product_space/positive.lean | 8 +- .../inner_product_space/projection.lean | 37 +++-- .../inner_product_space/rayleigh.lean | 6 +- .../inner_product_space/spectrum.lean | 2 +- .../inner_product_space/symmetric.lean | 12 +- src/analysis/inner_product_space/two_dim.lean | 14 +- src/analysis/matrix.lean | 2 +- src/analysis/quaternion.lean | 7 +- src/analysis/von_neumann_algebra/basic.lean | 5 +- .../euclidean/angle/oriented/affine.lean | 9 +- .../euclidean/angle/oriented/basic.lean | 4 +- .../euclidean/angle/oriented/right_angle.lean | 7 +- .../euclidean/angle/oriented/rotation.lean | 4 +- src/geometry/euclidean/angle/sphere.lean | 17 +- .../euclidean/angle/unoriented/affine.lean | 7 +- .../euclidean/angle/unoriented/basic.lean | 6 +- .../euclidean/angle/unoriented/conformal.lean | 12 +- .../angle/unoriented/right_angle.lean | 6 +- src/geometry/euclidean/basic.lean | 12 +- src/geometry/euclidean/circumcenter.lean | 12 +- src/geometry/euclidean/inversion.lean | 3 +- src/geometry/euclidean/monge_point.lean | 12 +- src/geometry/euclidean/sphere/basic.lean | 3 +- src/geometry/euclidean/sphere/power.lean | 2 +- src/geometry/euclidean/sphere/ptolemy.lean | 2 +- .../euclidean/sphere/second_inter.lean | 3 +- src/geometry/euclidean/triangle.lean | 6 +- src/geometry/manifold/instances/sphere.lean | 10 +- src/linear_algebra/matrix/ldl.lean | 16 +- src/linear_algebra/matrix/pos_def.lean | 14 +- src/linear_algebra/matrix/spectrum.lean | 4 +- .../function/ae_eq_of_integral.lean | 5 +- .../conditional_expectation/basic.lean | 27 ++-- src/measure_theory/function/l1_space.lean | 3 +- src/measure_theory/function/l2_space.lean | 2 +- src/measure_theory/function/lp_space.lean | 2 +- .../function/special_functions/inner.lean | 3 +- .../function/strongly_measurable/inner.lean | 5 +- src/measure_theory/integral/set_integral.lean | 8 +- src/measure_theory/measure/haar_of_basis.lean | 3 +- src/measure_theory/measure/haar_of_inner.lean | 5 +- 57 files changed, 367 insertions(+), 288 deletions(-) diff --git a/archive/100-theorems-list/57_herons_formula.lean b/archive/100-theorems-list/57_herons_formula.lean index fc62c68b265a5..7d288df537ac0 100644 --- a/archive/100-theorems-list/57_herons_formula.lean +++ b/archive/100-theorems-list/57_herons_formula.lean @@ -23,8 +23,8 @@ open_locale real euclidean_geometry local notation `√` := real.sqrt -variables {V : Type*} {P : Type*} [inner_product_space ℝ V] [metric_space P] - [normed_add_torsor V P] +variables {V : Type*} {P : Type*} + [normed_add_comm_group V] [inner_product_space ℝ V] [metric_space P] [normed_add_torsor V P] include V diff --git a/archive/imo/imo2019_q2.lean b/archive/imo/imo2019_q2.lean index d2af904f035d4..ca6781c82a7e5 100644 --- a/archive/imo/imo2019_q2.lean +++ b/archive/imo/imo2019_q2.lean @@ -63,7 +63,8 @@ open_locale affine euclidean_geometry real local attribute [instance] fact_finite_dimensional_of_finrank_eq_succ -variables (V : Type*) (Pt : Type*) [inner_product_space ℝ V] [metric_space Pt] +variables (V : Type*) (Pt : Type*) +variables [normed_add_comm_group V] [inner_product_space ℝ V] [metric_space Pt] variables [normed_add_torsor V Pt] [hd2 : fact (finrank ℝ V = 2)] include hd2 diff --git a/src/analysis/calculus/bump_function_inner.lean b/src/analysis/calculus/bump_function_inner.lean index 47c1955e56285..af3a1238fc7a4 100644 --- a/src/analysis/calculus/bump_function_inner.lean +++ b/src/analysis/calculus/bump_function_inner.lean @@ -306,7 +306,7 @@ nonempty.some hb.out /-- Any inner product space has smooth bump functions. -/ @[priority 100] instance has_cont_diff_bump_of_inner_product_space - (E : Type*) [inner_product_space ℝ E] : has_cont_diff_bump E := + (E : Type*) [normed_add_comm_group E] [inner_product_space ℝ E] : has_cont_diff_bump E := let e : cont_diff_bump_base E := { to_fun := λ R x, real.smooth_transition ((R - ‖x‖) / (R - 1)), mem_Icc := λ R x, ⟨real.smooth_transition.nonneg _, real.smooth_transition.le_one _⟩, @@ -331,7 +331,7 @@ let e : cont_diff_bump_base E := exact cont_diff_at_const.congr_of_eventually_eq this }, { refine real.smooth_transition.cont_diff_at.comp _ _, refine cont_diff_at.div _ _ (sub_pos.2 hR).ne', - { exact cont_diff_at_fst.sub (cont_diff_at_snd.norm hx) }, + { exact cont_diff_at_fst.sub (cont_diff_at_snd.norm ℝ hx) }, { exact cont_diff_at_fst.sub cont_diff_at_const } } end, eq_one := λ R hR x hx, real.smooth_transition.one_of_one_le $ diff --git a/src/analysis/calculus/conformal/inner_product.lean b/src/analysis/calculus/conformal/inner_product.lean index 7d84a6b327469..6daefbcfb7ebf 100644 --- a/src/analysis/calculus/conformal/inner_product.lean +++ b/src/analysis/calculus/conformal/inner_product.lean @@ -15,7 +15,9 @@ is conformal at `x` iff the derivative preserves inner products up to a scalar m noncomputable theory -variables {E F : Type*} [inner_product_space ℝ E] [inner_product_space ℝ F] +variables {E F : Type*} +variables [normed_add_comm_group E] [normed_add_comm_group F] +variables [inner_product_space ℝ E] [inner_product_space ℝ F] open_locale real_inner_product_space diff --git a/src/analysis/convex/cone/basic.lean b/src/analysis/convex/cone/basic.lean index e31d4ac163d73..ca62b0911010e 100644 --- a/src/analysis/convex/cone/basic.lean +++ b/src/analysis/convex/cone/basic.lean @@ -717,7 +717,7 @@ end /-! ### The dual cone -/ section dual -variables {H : Type*} [inner_product_space ℝ H] (s t : set H) +variables {H : Type*} [normed_add_comm_group H] [inner_product_space ℝ H] (s t : set H) open_locale real_inner_product_space /-- The dual cone is the cone consisting of all points `y` such that for diff --git a/src/analysis/fourier/add_circle.lean b/src/analysis/fourier/add_circle.lean index 4ad7a556d91e6..944f5424ac693 100644 --- a/src/analysis/fourier/add_circle.lean +++ b/src/analysis/fourier/add_circle.lean @@ -426,7 +426,7 @@ begin { exact_mod_cast lp.norm_rpow_eq_tsum _ (fourier_basis.repr f), norm_num }, have H₂ : ‖fourier_basis.repr f‖ ^ 2 = ‖f‖ ^ 2 := by simp, - have H₃ := congr_arg is_R_or_C.re (@L2.inner_def (add_circle T) ℂ ℂ _ _ _ _ f f), + have H₃ := congr_arg is_R_or_C.re (@L2.inner_def (add_circle T) ℂ ℂ _ _ _ _ _ f f), rw ← integral_re at H₃, { simp only [← norm_sq_eq_inner] at H₃, rw [← H₁, H₂, H₃], }, diff --git a/src/analysis/inner_product_space/adjoint.lean b/src/analysis/inner_product_space/adjoint.lean index 12ffdcc25e197..294543aa89761 100644 --- a/src/analysis/inner_product_space/adjoint.lean +++ b/src/analysis/inner_product_space/adjoint.lean @@ -44,6 +44,7 @@ open is_R_or_C open_locale complex_conjugate variables {𝕜 E F G : Type*} [is_R_or_C 𝕜] +variables [normed_add_comm_group E] [normed_add_comm_group F] [normed_add_comm_group G] variables [inner_product_space 𝕜 E] [inner_product_space 𝕜 F] [inner_product_space 𝕜 G] local notation `⟪`x`, `y`⟫` := @inner 𝕜 _ _ x y @@ -206,7 +207,9 @@ end⟩ section real -variables {E' : Type*} {F' : Type*} [inner_product_space ℝ E'] [inner_product_space ℝ F'] +variables {E' : Type*} {F' : Type*} +variables [normed_add_comm_group E'] [normed_add_comm_group F'] +variables [inner_product_space ℝ E'] [inner_product_space ℝ F'] variables [complete_space E'] [complete_space F'] -- Todo: Generalize this to `is_R_or_C`. @@ -402,7 +405,9 @@ by { rw [is_self_adjoint_iff', is_symmetric, ← linear_map.eq_adjoint_iff], exa section real -variables {E' : Type*} {F' : Type*} [inner_product_space ℝ E'] [inner_product_space ℝ F'] +variables {E' : Type*} {F' : Type*} +variables [normed_add_comm_group E'] [normed_add_comm_group F'] +variables [inner_product_space ℝ E'] [inner_product_space ℝ F'] variables [finite_dimensional ℝ E'] [finite_dimensional ℝ F'] -- Todo: Generalize this to `is_R_or_C`. diff --git a/src/analysis/inner_product_space/basic.lean b/src/analysis/inner_product_space/basic.lean index 1ba001bd04c70..4a82db4b3094e 100644 --- a/src/analysis/inner_product_space/basic.lean +++ b/src/analysis/inner_product_space/basic.lean @@ -14,8 +14,8 @@ import linear_algebra.bilinear_form # Inner product space This file defines inner product spaces and proves the basic properties. We do not formally -define Hilbert spaces, but they can be obtained using the pair of assumptions -`[inner_product_space 𝕜 E] [complete_space E]`. +define Hilbert spaces, but they can be obtained using the set of assumptions +`[normed_add_comm_group E] [inner_product_space 𝕜 E] [complete_space E]`. An inner product space is a vector space endowed with an inner product. It generalizes the notion of dot product in `ℝ^n` and provides the means of defining the length of a vector and the angle between @@ -99,16 +99,13 @@ spaces. To construct a norm from an inner product, see `inner_product_space.of_core`. -/ -class inner_product_space (𝕜 : Type*) (E : Type*) [is_R_or_C 𝕜] - extends normed_add_comm_group E, normed_space 𝕜 E, has_inner 𝕜 E := +class inner_product_space (𝕜 : Type*) (E : Type*) [is_R_or_C 𝕜] [normed_add_comm_group E] + extends normed_space 𝕜 E, has_inner 𝕜 E := (norm_sq_eq_inner : ∀ (x : E), ‖x‖^2 = re (inner x x)) (conj_symm : ∀ x y, conj (inner y x) = inner x y) (add_left : ∀ x y z, inner (x + y) z = inner x z + inner y z) (smul_left : ∀ x y r, inner (r • x) y = (conj r) * inner x y) -attribute [nolint dangerous_instance] inner_product_space.to_normed_add_comm_group --- note [is_R_or_C instance] - /-! ### Constructing a normed space structure from an inner product @@ -361,13 +358,15 @@ def to_normed_space : normed_space 𝕜 F := end inner_product_space.of_core +section +local attribute [instance] inner_product_space.of_core.to_normed_add_comm_group + /-- Given a `inner_product_space.core` structure on a space, one can use it to turn -the space into an inner product space, constructing the norm out of the inner product -/ +the space into an inner product space. The `normed_add_comm_group` structure is expected +to already be defined with `inner_product_space.of_core.to_normed_add_comm_group`. -/ def inner_product_space.of_core [add_comm_group F] [module 𝕜 F] (c : inner_product_space.core 𝕜 F) : inner_product_space 𝕜 F := begin - letI : normed_add_comm_group F := - @inner_product_space.of_core.to_normed_add_comm_group 𝕜 F _ _ _ c, letI : normed_space 𝕜 F := @inner_product_space.of_core.to_normed_space 𝕜 F _ _ _ c, exact { norm_sq_eq_inner := λ x, begin @@ -378,9 +377,12 @@ begin ..c } end +end + /-! ### Properties of inner product spaces -/ -variables [inner_product_space 𝕜 E] [inner_product_space ℝ F] +variables [normed_add_comm_group E] [inner_product_space 𝕜 E] +variables [normed_add_comm_group F] [inner_product_space ℝ F] variables [dec_E : decidable_eq E] local notation `⟪`x`, `y`⟫` := @inner 𝕜 _ _ x y local notation `IK` := @is_R_or_C.I 𝕜 _ @@ -393,7 +395,7 @@ export inner_product_space (norm_sq_eq_inner) section basic_properties @[simp] lemma inner_conj_symm (x y : E) : ⟪y, x⟫† = ⟪x, y⟫ := inner_product_space.conj_symm _ _ -lemma real_inner_comm (x y : F) : ⟪y, x⟫_ℝ = ⟪x, y⟫_ℝ := @inner_conj_symm ℝ _ _ _ x y +lemma real_inner_comm (x y : F) : ⟪y, x⟫_ℝ = ⟪x, y⟫_ℝ := @inner_conj_symm ℝ _ _ _ _ x y lemma inner_eq_zero_symm {x y : E} : ⟪x, y⟫ = 0 ↔ ⟪y, x⟫ = 0 := ⟨λ h, by simp [←inner_conj_symm, h], λ h, by simp [←inner_conj_symm, h]⟩ @@ -498,7 +500,7 @@ by simp only [inner_zero_right, add_monoid_hom.map_zero] lemma inner_self_nonneg {x : E} : 0 ≤ re ⟪x, x⟫ := by rw [←norm_sq_eq_inner]; exact pow_nonneg (norm_nonneg x) 2 -lemma real_inner_self_nonneg {x : F} : 0 ≤ ⟪x, x⟫_ℝ := @inner_self_nonneg ℝ F _ _ x +lemma real_inner_self_nonneg {x : F} : 0 ≤ ⟪x, x⟫_ℝ := @inner_self_nonneg ℝ F _ _ _ x @[simp] lemma inner_self_eq_zero {x : E} : ⟪x, x⟫ = 0 ↔ x = 0 := begin @@ -520,7 +522,7 @@ inner_self_eq_zero.not begin split, { intro h, - rw ←inner_self_eq_zero, + rw ←@inner_self_eq_zero 𝕜, have H₁ : re ⟪x, x⟫ ≥ 0, exact inner_self_nonneg, have H₂ : re ⟪x, x⟫ = 0, exact le_antisymm h H₁, rw is_R_or_C.ext_iff, @@ -530,7 +532,7 @@ begin end lemma real_inner_self_nonpos {x : F} : ⟪x, x⟫_ℝ ≤ 0 ↔ x = 0 := -by { have h := @inner_self_nonpos ℝ F _ _ x, simpa using h } +by { have h := @inner_self_nonpos ℝ F _ _ _ x, simpa using h } @[simp] lemma inner_self_re_to_K (x : E) : (re ⟪x, x⟫ : 𝕜) = ⟪x, x⟫ := is_R_or_C.ext_iff.2 ⟨by simp only [of_real_re], by simp only [inner_self_nonneg_im, of_real_im]⟩ @@ -553,7 +555,7 @@ lemma inner_self_abs_to_K (x : E) : (absK ⟪x, x⟫ : 𝕜) = ⟪x, x⟫ := by { rw [←inner_self_re_abs], exact inner_self_re_to_K _ } lemma real_inner_self_abs (x : F) : absR ⟪x, x⟫_ℝ = ⟪x, x⟫_ℝ := -by { have h := @inner_self_abs_to_K ℝ F _ _ x, simpa using h } +by { have h := @inner_self_abs_to_K ℝ F _ _ _ x, simpa using h } lemma inner_abs_conj_symm (x y : E) : abs ⟪x, y⟫ = abs ⟪y, x⟫ := by rw [←inner_conj_symm, abs_conj] @@ -606,10 +608,10 @@ variable (𝕜) include 𝕜 lemma ext_inner_left {x y : E} (h : ∀ v, ⟪v, x⟫ = ⟪v, y⟫) : x = y := -by rw [←sub_eq_zero, ←inner_self_eq_zero, inner_sub_right, sub_eq_zero, h (x - y)] +by rw [←sub_eq_zero, ←@inner_self_eq_zero 𝕜, inner_sub_right, sub_eq_zero, h (x - y)] lemma ext_inner_right {x y : E} (h : ∀ v, ⟪x, v⟫ = ⟪y, v⟫) : x = y := -by rw [←sub_eq_zero, ←inner_self_eq_zero, inner_sub_left, sub_eq_zero, h (x - y)] +by rw [←sub_eq_zero, ←@inner_self_eq_zero 𝕜, inner_sub_left, sub_eq_zero, h (x - y)] omit 𝕜 variable {𝕜} @@ -673,7 +675,7 @@ end lemma real_inner_mul_inner_self_le (x y : F) : ⟪x, y⟫_ℝ * ⟪x, y⟫_ℝ ≤ ⟪x, x⟫_ℝ * ⟪y, y⟫_ℝ := begin have h₁ : ⟪y, x⟫_ℝ = ⟪x, y⟫_ℝ := by rw [←inner_conj_symm]; refl, - have h₂ := @inner_mul_inner_self_le ℝ F _ _ x y, + have h₂ := @inner_mul_inner_self_le ℝ F _ _ _ x y, dsimp at h₂, have h₃ := abs_mul_abs_self ⟪x, y⟫_ℝ, rw [h₁] at h₂, @@ -727,7 +729,7 @@ begin { intros h, split, { intros i, - have h' : ‖v i‖ ^ 2 = 1 ^ 2 := by simp [norm_sq_eq_inner, h i i], + have h' : ‖v i‖ ^ 2 = 1 ^ 2 := by simp [@norm_sq_eq_inner 𝕜, h i i], have h₁ : 0 ≤ ‖v i‖ := norm_nonneg _, have h₂ : (0:ℝ) ≤ 1 := zero_le_one, rwa sq_eq_sq h₁ h₂ at h' }, @@ -964,25 +966,26 @@ calc ‖x‖ = sqrt (‖x‖ ^ 2) : (sqrt_sq (norm_nonneg _)).symm ... = sqrt (re ⟪x, x⟫) : congr_arg _ (norm_sq_eq_inner _) lemma norm_eq_sqrt_real_inner (x : F) : ‖x‖ = sqrt ⟪x, x⟫_ℝ := -by { have h := @norm_eq_sqrt_inner ℝ F _ _ x, simpa using h } +by { have h := @norm_eq_sqrt_inner ℝ F _ _ _ x, simpa using h } lemma inner_self_eq_norm_mul_norm (x : E) : re ⟪x, x⟫ = ‖x‖ * ‖x‖ := -by rw [norm_eq_sqrt_inner, ←sqrt_mul inner_self_nonneg (re ⟪x, x⟫), +by rw [@norm_eq_sqrt_inner 𝕜, ←sqrt_mul inner_self_nonneg (re ⟪x, x⟫), sqrt_mul_self inner_self_nonneg] lemma inner_self_eq_norm_sq (x : E) : re ⟪x, x⟫ = ‖x‖^2 := by rw [pow_two, inner_self_eq_norm_mul_norm] lemma real_inner_self_eq_norm_mul_norm (x : F) : ⟪x, x⟫_ℝ = ‖x‖ * ‖x‖ := -by { have h := @inner_self_eq_norm_mul_norm ℝ F _ _ x, simpa using h } +by { have h := @inner_self_eq_norm_mul_norm ℝ F _ _ _ x, simpa using h } lemma real_inner_self_eq_norm_sq (x : F) : ⟪x, x⟫_ℝ = ‖x‖^2 := by rw [pow_two, real_inner_self_eq_norm_mul_norm] +variables (𝕜) /-- Expand the square -/ lemma norm_add_sq (x y : E) : ‖x + y‖^2 = ‖x‖^2 + 2 * (re ⟪x, y⟫) + ‖y‖^2 := begin - repeat {rw [sq, ←inner_self_eq_norm_mul_norm]}, + repeat {rw [sq, ←@inner_self_eq_norm_mul_norm 𝕜]}, rw [inner_add_add_self, two_mul], simp only [add_assoc, add_left_inj, add_right_inj, add_monoid_hom.map_add], rw [←inner_conj_symm, conj_re], @@ -992,7 +995,7 @@ alias norm_add_sq ← norm_add_pow_two /-- Expand the square -/ lemma norm_add_sq_real (x y : F) : ‖x + y‖^2 = ‖x‖^2 + 2 * ⟪x, y⟫_ℝ + ‖y‖^2 := -by { have h := @norm_add_sq ℝ _ _ _ x y, simpa using h } +by { have h := @norm_add_sq ℝ _ _ _ _ x y, simpa using h } alias norm_add_sq_real ← norm_add_pow_two_real @@ -1002,12 +1005,12 @@ by { repeat {rw [← sq]}, exact norm_add_sq _ _ } /-- Expand the square -/ lemma norm_add_mul_self_real (x y : F) : ‖x + y‖ * ‖x + y‖ = ‖x‖ * ‖x‖ + 2 * ⟪x, y⟫_ℝ + ‖y‖ * ‖y‖ := -by { have h := @norm_add_mul_self ℝ _ _ _ x y, simpa using h } +by { have h := @norm_add_mul_self ℝ _ _ _ _ x y, simpa using h } /-- Expand the square -/ lemma norm_sub_sq (x y : E) : ‖x - y‖^2 = ‖x‖^2 - 2 * (re ⟪x, y⟫) + ‖y‖^2 := begin - repeat {rw [sq, ←inner_self_eq_norm_mul_norm]}, + repeat {rw [sq, ←@inner_self_eq_norm_mul_norm 𝕜]}, rw [inner_sub_sub_self], calc re (⟪x, x⟫ - ⟪x, y⟫ - ⟪y, x⟫ + ⟪y, y⟫) @@ -1022,7 +1025,7 @@ alias norm_sub_sq ← norm_sub_pow_two /-- Expand the square -/ lemma norm_sub_sq_real (x y : F) : ‖x - y‖^2 = ‖x‖^2 - 2 * ⟪x, y⟫_ℝ + ‖y‖^2 := -norm_sub_sq _ _ +@norm_sub_sq ℝ _ _ _ _ _ _ alias norm_sub_sq_real ← norm_sub_pow_two_real @@ -1032,7 +1035,7 @@ by { repeat {rw [← sq]}, exact norm_sub_sq _ _ } /-- Expand the square -/ lemma norm_sub_mul_self_real (x y : F) : ‖x - y‖ * ‖x - y‖ = ‖x‖ * ‖x‖ - 2 * ⟪x, y⟫_ℝ + ‖y‖ * ‖y‖ := -by { have h := @norm_sub_mul_self ℝ _ _ _ x y, simpa using h } +by { have h := @norm_sub_mul_self ℝ _ _ _ _ x y, simpa using h } /-- Cauchy–Schwarz inequality with norm -/ lemma abs_inner_le_norm (x y : E) : abs ⟪x, y⟫ ≤ ‖x‖ * ‖y‖ := @@ -1056,46 +1059,47 @@ le_trans (re_le_abs (inner x y)) (abs_inner_le_norm x y) /-- Cauchy–Schwarz inequality with norm -/ lemma abs_real_inner_le_norm (x y : F) : absR ⟪x, y⟫_ℝ ≤ ‖x‖ * ‖y‖ := -by { have h := @abs_inner_le_norm ℝ F _ _ x y, simpa using h } +by { have h := @abs_inner_le_norm ℝ F _ _ _ x y, simpa using h } /-- Cauchy–Schwarz inequality with norm -/ lemma real_inner_le_norm (x y : F) : ⟪x, y⟫_ℝ ≤ ‖x‖ * ‖y‖ := le_trans (le_abs_self _) (abs_real_inner_le_norm _ _) include 𝕜 +variables (𝕜) lemma parallelogram_law_with_norm (x y : E) : ‖x + y‖ * ‖x + y‖ + ‖x - y‖ * ‖x - y‖ = 2 * (‖x‖ * ‖x‖ + ‖y‖ * ‖y‖) := begin - simp only [← inner_self_eq_norm_mul_norm], + simp only [← @inner_self_eq_norm_mul_norm 𝕜], rw [← re.map_add, parallelogram_law, two_mul, two_mul], simp only [re.map_add], end lemma parallelogram_law_with_nnnorm (x y : E) : ‖x + y‖₊ * ‖x + y‖₊ + ‖x - y‖₊ * ‖x - y‖₊ = 2 * (‖x‖₊ * ‖x‖₊ + ‖y‖₊ * ‖y‖₊) := -subtype.ext $ parallelogram_law_with_norm x y - +subtype.ext $ parallelogram_law_with_norm 𝕜 x y +variables {𝕜} omit 𝕜 /-- Polarization identity: The real part of the inner product, in terms of the norm. -/ lemma re_inner_eq_norm_add_mul_self_sub_norm_mul_self_sub_norm_mul_self_div_two (x y : E) : re ⟪x, y⟫ = (‖x + y‖ * ‖x + y‖ - ‖x‖ * ‖x‖ - ‖y‖ * ‖y‖) / 2 := -by { rw norm_add_mul_self, ring } +by { rw @norm_add_mul_self 𝕜, ring } /-- Polarization identity: The real part of the inner product, in terms of the norm. -/ lemma re_inner_eq_norm_mul_self_add_norm_mul_self_sub_norm_sub_mul_self_div_two (x y : E) : re ⟪x, y⟫ = (‖x‖ * ‖x‖ + ‖y‖ * ‖y‖ - ‖x - y‖ * ‖x - y‖) / 2 := -by { rw [norm_sub_mul_self], ring } +by { rw [@norm_sub_mul_self 𝕜], ring } /-- Polarization identity: The real part of the inner product, in terms of the norm. -/ lemma re_inner_eq_norm_add_mul_self_sub_norm_sub_mul_self_div_four (x y : E) : re ⟪x, y⟫ = (‖x + y‖ * ‖x + y‖ - ‖x - y‖ * ‖x - y‖) / 4 := -by { rw [norm_add_mul_self, norm_sub_mul_self], ring } +by { rw [@norm_add_mul_self 𝕜, @norm_sub_mul_self 𝕜], ring } /-- Polarization identity: The imaginary part of the inner product, in terms of the norm. -/ lemma im_inner_eq_norm_sub_I_smul_mul_self_sub_norm_add_I_smul_mul_self_div_four (x y : E) : im ⟪x, y⟫ = (‖x - IK • y‖ * ‖x - IK • y‖ - ‖x + IK • y‖ * ‖x + IK • y‖) / 4 := -by { simp only [norm_add_mul_self, norm_sub_mul_self, inner_smul_right, I_mul_re], ring } +by { simp only [@norm_add_mul_self 𝕜, @norm_sub_mul_self 𝕜, inner_smul_right, I_mul_re], ring } /-- Polarization identity: The inner product, in terms of the norm. -/ lemma inner_eq_sum_norm_sq_div_four (x y : E) : @@ -1132,15 +1136,14 @@ instance inner_product_space.to_uniform_convex_space : uniform_convex_space F := exact pow_pos hε _ }, rw sub_sub_cancel, refine le_sqrt_of_sq_le _, - rw [sq, eq_sub_iff_add_eq.2 (parallelogram_law_with_norm x y), ←sq (‖x - y‖), hx, hy], + rw [sq, eq_sub_iff_add_eq.2 (parallelogram_law_with_norm ℝ x y), ←sq (‖x - y‖), hx, hy], norm_num, exact pow_le_pow_of_le_left hε.le hxy _, end⟩ section complex -variables {V : Type*} -[inner_product_space ℂ V] +variables {V : Type*} [normed_add_comm_group V] [inner_product_space ℂ V] /-- A complex polarization identity, with a linear map @@ -1178,7 +1181,7 @@ begin split, { intro hT, ext x, - simp only [linear_map.zero_apply, ← inner_self_eq_zero, inner_map_polarization, hT], + simp only [linear_map.zero_apply, ← @inner_self_eq_zero ℂ, inner_map_polarization, hT], norm_num }, { rintro rfl x, simp only [linear_map.zero_apply, inner_zero_left] } @@ -1201,8 +1204,8 @@ end complex section variables {ι : Type*} {ι' : Type*} {ι'' : Type*} -variables {E' : Type*} [inner_product_space 𝕜 E'] -variables {E'' : Type*} [inner_product_space 𝕜 E''] +variables {E' : Type*} [normed_add_comm_group E'] [inner_product_space 𝕜 E'] +variables {E'' : Type*} [normed_add_comm_group E''] [inner_product_space 𝕜 E''] /-- A linear isometry preserves the inner product. -/ @[simp] lemma linear_isometry.inner_map_map (f : E →ₗᵢ[𝕜] E') (x y : E) : ⟪f x, f y⟫ = ⟪x, y⟫ := @@ -1215,7 +1218,7 @@ f.to_linear_isometry.inner_map_map x y /-- A linear map that preserves the inner product is a linear isometry. -/ def linear_map.isometry_of_inner (f : E →ₗ[𝕜] E') (h : ∀ x y, ⟪f x, f y⟫ = ⟪x, y⟫) : E →ₗᵢ[𝕜] E' := -⟨f, λ x, by simp only [norm_eq_sqrt_inner, h]⟩ +⟨f, λ x, by simp only [@norm_eq_sqrt_inner 𝕜, h]⟩ @[simp] lemma linear_map.coe_isometry_of_inner (f : E →ₗ[𝕜] E') (h) : ⇑(f.isometry_of_inner h) = f := rfl @@ -1355,7 +1358,7 @@ re_to_real.symm.trans $ lemma norm_add_sq_eq_norm_sq_add_norm_sq_iff_real_inner_eq_zero (x y : F) : ‖x + y‖ * ‖x + y‖ = ‖x‖ * ‖x‖ + ‖y‖ * ‖y‖ ↔ ⟪x, y⟫_ℝ = 0 := begin - rw [norm_add_mul_self, add_right_cancel_iff, add_right_eq_self, mul_eq_zero], + rw [@norm_add_mul_self ℝ, add_right_cancel_iff, add_right_eq_self, mul_eq_zero], norm_num end @@ -1369,7 +1372,7 @@ by rw [←norm_add_sq_eq_norm_sq_add_norm_sq_iff_real_inner_eq_zero, eq_comm, lemma norm_add_sq_eq_norm_sq_add_norm_sq_of_inner_eq_zero (x y : E) (h : ⟪x, y⟫ = 0) : ‖x + y‖ * ‖x + y‖ = ‖x‖ * ‖x‖ + ‖y‖ * ‖y‖ := begin - rw [norm_add_mul_self, add_right_cancel_iff, add_right_eq_self, mul_eq_zero], + rw [@norm_add_mul_self 𝕜, add_right_cancel_iff, add_right_eq_self, mul_eq_zero], apply or.inr, simp only [h, zero_re'], end @@ -1384,7 +1387,7 @@ inner product form. -/ lemma norm_sub_sq_eq_norm_sq_add_norm_sq_iff_real_inner_eq_zero (x y : F) : ‖x - y‖ * ‖x - y‖ = ‖x‖ * ‖x‖ + ‖y‖ * ‖y‖ ↔ ⟪x, y⟫_ℝ = 0 := begin - rw [norm_sub_mul_self, add_right_cancel_iff, sub_eq_add_neg, add_right_eq_self, neg_eq_zero, + rw [@norm_sub_mul_self ℝ, add_right_cancel_iff, sub_eq_add_neg, add_right_eq_self, neg_eq_zero, mul_eq_zero], norm_num end @@ -1407,7 +1410,7 @@ if they have the same norm. -/ lemma real_inner_add_sub_eq_zero_iff (x y : F) : ⟪x + y, x - y⟫_ℝ = 0 ↔ ‖x‖ = ‖y‖ := begin conv_rhs { rw ←mul_self_inj_of_nonneg (norm_nonneg _) (norm_nonneg _) }, - simp only [←inner_self_eq_norm_mul_norm, inner_add_left, inner_sub_right, + simp only [←@inner_self_eq_norm_mul_norm ℝ, inner_add_left, inner_sub_right, real_inner_comm y x, sub_eq_zero, re_to_real], split, { intro h, @@ -1421,7 +1424,8 @@ end lemma norm_sub_eq_norm_add {v w : E} (h : ⟪v, w⟫ = 0) : ‖w - v‖ = ‖w + v‖ := begin rw ←mul_self_inj_of_nonneg (norm_nonneg _) (norm_nonneg _), - simp only [h, ←inner_self_eq_norm_mul_norm, sub_neg_eq_add, sub_zero, map_sub, zero_re', zero_sub, + simp only [h, ←@inner_self_eq_norm_mul_norm 𝕜, sub_neg_eq_add, sub_zero, map_sub, zero_re', + zero_sub, add_zero, map_add, inner_add_right, inner_sub_left, inner_sub_right, inner_re_symm, zero_add] end @@ -1513,8 +1517,8 @@ begin have ht0 : ⟪x, t⟫ = 0, { rw [ht, inner_sub_right, inner_smul_right, hr], norm_cast, - rw [←inner_self_eq_norm_mul_norm, inner_self_re_to_K, - div_mul_cancel _ (λ h, hx0 (inner_self_eq_zero.1 h)), sub_self] }, + rw [←@inner_self_eq_norm_mul_norm 𝕜, inner_self_re_to_K, + div_mul_cancel _ (λ h, hx0 ((@inner_self_eq_zero 𝕜 _ _ _ _ _).1 h)), sub_self] }, replace h : ‖r • x‖ / ‖t + r • x‖ = 1, { rw [←sub_add_cancel y (r • x), ←ht, inner_add_right, ht0, zero_add, inner_smul_right, is_R_or_C.abs_div, is_R_or_C.abs_mul, ←inner_self_re_abs, @@ -1531,7 +1535,7 @@ begin have h2 : ‖r • x‖ ^ 2 = ‖t + r • x‖ ^ 2, { rw [eq_of_div_eq_one h] }, replace h2 : ⟪r • x, r • x⟫ = ⟪t, t⟫ + ⟪t, r • x⟫ + ⟪r • x, t⟫ + ⟪r • x, r • x⟫, - { rw [sq, sq, ←inner_self_eq_norm_mul_norm, ←inner_self_eq_norm_mul_norm ] at h2, + { rw [sq, sq, ←@inner_self_eq_norm_mul_norm 𝕜, ←@inner_self_eq_norm_mul_norm 𝕜] at h2, have h2' := congr_arg (λ z : ℝ, (z : 𝕜)) h2, simp_rw [inner_self_re_to_K, inner_add_add_self] at h2', exact h2' }, @@ -1555,7 +1559,7 @@ a multiple of the other. One form of equality case for Cauchy-Schwarz. -/ lemma abs_real_inner_div_norm_mul_norm_eq_one_iff (x y : F) : absR (⟪x, y⟫_ℝ / (‖x‖ * ‖y‖)) = 1 ↔ (x ≠ 0 ∧ ∃ (r : ℝ), r ≠ 0 ∧ y = r • x) := begin - have := @abs_inner_div_norm_mul_norm_eq_one_iff ℝ F _ _ x y, + have := @abs_inner_div_norm_mul_norm_eq_one_iff ℝ F _ _ _ x y, simpa [coe_real_eq_id] using this, end @@ -1648,7 +1652,7 @@ begin by simp [h, show (2:ℝ) ≠ 0, by norm_num, sub_eq_zero] ... ↔ ‖(‖y‖:𝕜) • x - (‖x‖:𝕜) • y‖ * ‖(‖y‖:𝕜) • x - (‖x‖:𝕜) • y‖ = 0 : begin - simp only [norm_sub_mul_self, inner_smul_left, inner_smul_right, norm_smul, conj_of_real, + simp only [@norm_sub_mul_self 𝕜, inner_smul_left, inner_smul_right, norm_smul, conj_of_real, is_R_or_C.norm_eq_abs, abs_of_real, of_real_im, of_real_re, mul_re, abs_norm_eq_norm], refine eq.congr _ rfl, ring @@ -1745,7 +1749,7 @@ variables {𝕜} namespace continuous_linear_map -variables {E' : Type*} [inner_product_space 𝕜 E'] +variables {E' : Type*} [normed_add_comm_group E'] [inner_product_space 𝕜 E'] /-- Given `f : E →L[𝕜] E'`, construct the continuous sesquilinear form `λ x y, ⟪x, A y⟫`, given as a continuous linear map. -/ @@ -1762,7 +1766,7 @@ begin refine op_norm_le_bound _ (mul_nonneg (norm_nonneg _) (norm_nonneg _)) _, intro x, have h₁ : ‖f x‖ ≤ ‖f‖ * ‖x‖ := le_op_norm _ _, - have h₂ := @norm_inner_le_norm 𝕜 E' _ _ v (f x), + have h₂ := @norm_inner_le_norm 𝕜 E' _ _ _ v (f x), calc ‖⟪v, f x⟫‖ ≤ ‖v‖ * ‖f x‖ : h₂ ... ≤ ‖v‖ * (‖f‖ * ‖x‖) : mul_le_mul_of_nonneg_left h₁ (norm_nonneg v) ... = ‖f‖ * ‖v‖ * ‖x‖ : by ring, @@ -1809,8 +1813,8 @@ begin suffices hbf: ‖x - ∑ i in s, ⟪v i, x⟫ • (v i)‖ ^ 2 = ‖x‖ ^ 2 - ∑ i in s, ‖⟪v i, x⟫‖ ^ 2, { rw [←sub_nonneg, ←hbf], simp only [norm_nonneg, pow_nonneg], }, - rw [norm_sub_sq, sub_add], - simp only [inner_product_space.norm_sq_eq_inner, inner_sum], + rw [@norm_sub_sq 𝕜, sub_add], + simp only [@inner_product_space.norm_sq_eq_inner 𝕜, inner_sum], simp only [sum_inner, two_mul, inner_smul_right, inner_conj_symm, ←mul_assoc, h₂, ←h₃, inner_conj_symm, add_monoid_hom.map_sum, finset.mul_sum, ←finset.sum_sub_distrib, inner_smul_left, add_sub_cancel'], @@ -1841,8 +1845,7 @@ end bessels_inequality /-- A field `𝕜` satisfying `is_R_or_C` is itself a `𝕜`-inner product space. -/ instance is_R_or_C.inner_product_space : inner_product_space 𝕜 𝕜 := -{ to_normed_add_comm_group := non_unital_normed_ring.to_normed_add_comm_group, - inner := λ x y, conj x * y, +{ inner := λ x y, conj x * y, norm_sq_eq_inner := λ x, by { unfold inner, rw [mul_comm, mul_conj, of_real_re, norm_sq_eq_def'] }, conj_symm := λ x y, by simp only [mul_comm, map_mul, star_ring_end_self_apply], @@ -1855,10 +1858,9 @@ instance is_R_or_C.inner_product_space : inner_product_space 𝕜 𝕜 := /-- Induced inner product on a submodule. -/ instance submodule.inner_product_space (W : submodule 𝕜 E) : inner_product_space 𝕜 W := -{ to_normed_add_comm_group := submodule.normed_add_comm_group _, - inner := λ x y, ⟪(x:E), (y:E)⟫, - conj_symm := λ _ _, inner_conj_symm _ _ , - norm_sq_eq_inner := λ _, norm_sq_eq_inner _, +{ inner := λ x y, ⟪(x:E), (y:E)⟫, + conj_symm := λ _ _, inner_conj_symm _ _, + norm_sq_eq_inner := λ x, norm_sq_eq_inner (x : E), add_left := λ _ _ _, inner_add_left _ _ _, smul_left := λ _ _ _, inner_smul_left _ _ _, ..submodule.normed_space W } @@ -1868,11 +1870,11 @@ instance submodule.inner_product_space (W : submodule 𝕜 E) : inner_product_sp lemma orthonormal.cod_restrict {ι : Type*} {v : ι → E} (hv : orthonormal 𝕜 v) (s : submodule 𝕜 E) (hvs : ∀ i, v i ∈ s) : - @orthonormal 𝕜 s _ _ ι (set.cod_restrict v s hvs) := + @orthonormal 𝕜 s _ _ _ ι (set.cod_restrict v s hvs) := s.subtypeₗᵢ.orthonormal_comp_iff.mp hv lemma orthonormal_span {ι : Type*} {v : ι → E} (hv : orthonormal 𝕜 v) : - @orthonormal 𝕜 (submodule.span 𝕜 (set.range v)) _ _ ι + @orthonormal 𝕜 (submodule.span 𝕜 (set.range v)) _ _ _ ι (λ i : ι, ⟨v i, submodule.subset_span (set.mem_range_self i)⟩) := hv.cod_restrict (submodule.span 𝕜 (set.range v)) (λ i, submodule.subset_span (set.mem_range_self i)) @@ -1894,11 +1896,13 @@ product space structure on each of the submodules is important -- for example, w their Hilbert sum (`pi_lp V 2`). For example, given an orthonormal set of vectors `v : ι → E`, we have an associated orthogonal family of one-dimensional subspaces of `E`, which it is convenient to be able to discuss using `ι → 𝕜` rather than `Π i : ι, span 𝕜 (v i)`. -/ -def orthogonal_family (G : ι → Type*) [Π i, inner_product_space 𝕜 (G i)] (V : Π i, G i →ₗᵢ[𝕜] E) : +def orthogonal_family (G : ι → Type*) + [Π i, normed_add_comm_group (G i)] [Π i, inner_product_space 𝕜 (G i)] (V : Π i, G i →ₗᵢ[𝕜] E) : Prop := ∀ ⦃i j⦄, i ≠ j → ∀ v : G i, ∀ w : G j, ⟪V i v, V j w⟫ = 0 -variables {𝕜} {G : ι → Type*} [Π i, inner_product_space 𝕜 (G i)] {V : Π i, G i →ₗᵢ[𝕜] E} +variables {𝕜} {G : ι → Type*} + [Π i, normed_add_comm_group (G i)] [Π i, inner_product_space 𝕜 (G i)] {V : Π i, G i →ₗᵢ[𝕜] E} (hV : orthogonal_family 𝕜 G V) [dec_V : Π i (x : G i), decidable (x ≠ 0)] lemma orthonormal.orthogonal_family {v : ι → E} (hv : orthonormal 𝕜 v) : @@ -2111,8 +2115,7 @@ registered as an instance since it creates problems with the case `𝕜 = ℝ`, proof to obtain a real inner product space structure from a given `𝕜`-inner product space structure. -/ def inner_product_space.is_R_or_C_to_real : inner_product_space ℝ E := -{ to_normed_add_comm_group := inner_product_space.to_normed_add_comm_group 𝕜, - norm_sq_eq_inner := norm_sq_eq_inner, +{ norm_sq_eq_inner := norm_sq_eq_inner, conj_symm := λ x y, inner_re_symm _ _, add_left := λ x y z, by { change re ⟪x + y, z⟫ = re ⟪x, z⟫ + re ⟪y, z⟫, @@ -2135,14 +2138,16 @@ by simp [real_inner_eq_re_inner, inner_smul_right] omit 𝕜 /-- A complex inner product implies a real inner product -/ -instance inner_product_space.complex_to_real [inner_product_space ℂ G] : inner_product_space ℝ G := +instance inner_product_space.complex_to_real + [normed_add_comm_group G] [inner_product_space ℂ G] : inner_product_space ℝ G := inner_product_space.is_R_or_C_to_real ℂ G @[simp] protected lemma complex.inner (w z : ℂ) : ⟪w, z⟫_ℝ = (conj w * z).re := rfl /-- The inner product on an inner product space of dimension 2 can be evaluated in terms of a complex-number representation of the space. -/ -lemma inner_map_complex [inner_product_space ℝ G] (f : G ≃ₗᵢ[ℝ] ℂ) (x y : G) : +lemma inner_map_complex [normed_add_comm_group G] [inner_product_space ℝ G] + (f : G ≃ₗᵢ[ℝ] ℂ) (x y : G) : ⟪x, y⟫_ℝ = (conj (f x) * f y).re := by rw [← complex.inner, f.inner_map_map] @@ -2417,8 +2422,7 @@ protected lemma continuous.inner {α : Type*} [topological_space α] uniform_space.completion.continuous_inner.comp (hf.prod_mk hg : _) instance : inner_product_space 𝕜 (completion E) := -{ to_normed_add_comm_group := infer_instance, - norm_sq_eq_inner := λ x, completion.induction_on x +{ norm_sq_eq_inner := λ x, completion.induction_on x (is_closed_eq (continuous_norm.pow 2) (continuous_re.comp (continuous.inner continuous_id' continuous_id'))) diff --git a/src/analysis/inner_product_space/calculus.lean b/src/analysis/inner_product_space/calculus.lean index d146fb60fc3bc..4abc6d715a9e2 100644 --- a/src/analysis/inner_product_space/calculus.lean +++ b/src/analysis/inner_product_space/calculus.lean @@ -31,16 +31,17 @@ open_locale big_operators classical topology section deriv_inner variables {𝕜 E F : Type*} [is_R_or_C 𝕜] -variables [inner_product_space 𝕜 E] [inner_product_space ℝ F] +variables [normed_add_comm_group E] [inner_product_space 𝕜 E] +variables [normed_add_comm_group F] [inner_product_space ℝ F] local notation `⟪`x`, `y`⟫` := @inner 𝕜 _ _ x y -variables [normed_space ℝ E] +variables (𝕜) [normed_space ℝ E] /-- Derivative of the inner product. -/ def fderiv_inner_clm (p : E × E) : E × E →L[ℝ] 𝕜 := is_bounded_bilinear_map_inner.deriv p @[simp] lemma fderiv_inner_clm_apply (p x : E × E) : - fderiv_inner_clm p x = ⟪p.1, x.2⟫ + ⟪x.1, p.2⟫ := rfl + fderiv_inner_clm 𝕜 p x = ⟪p.1, x.2⟫ + ⟪x.1, p.2⟫ := rfl lemma cont_diff_inner {n} : cont_diff ℝ n (λ p : E × E, ⟪p.1, p.2⟫) := is_bounded_bilinear_map_inner.cont_diff @@ -65,11 +66,11 @@ cont_diff_at_inner.comp_cont_diff_within_at x (hf.prod hg) lemma cont_diff_at.inner (hf : cont_diff_at ℝ n f x) (hg : cont_diff_at ℝ n g x) : cont_diff_at ℝ n (λ x, ⟪f x, g x⟫) x := -hf.inner hg +hf.inner 𝕜 hg lemma cont_diff_on.inner (hf : cont_diff_on ℝ n f s) (hg : cont_diff_on ℝ n g s) : cont_diff_on ℝ n (λ x, ⟪f x, g x⟫) s := -λ x hx, (hf x hx).inner (hg x hx) +λ x hx, (hf x hx).inner 𝕜 (hg x hx) lemma cont_diff.inner (hf : cont_diff ℝ n f) (hg : cont_diff ℝ n g) : cont_diff ℝ n (λ x, ⟪f x, g x⟫) := @@ -77,27 +78,27 @@ cont_diff_inner.comp (hf.prod hg) lemma has_fderiv_within_at.inner (hf : has_fderiv_within_at f f' s x) (hg : has_fderiv_within_at g g' s x) : - has_fderiv_within_at (λ t, ⟪f t, g t⟫) ((fderiv_inner_clm (f x, g x)).comp $ f'.prod g') s x := + has_fderiv_within_at (λ t, ⟪f t, g t⟫) ((fderiv_inner_clm 𝕜 (f x, g x)).comp $ f'.prod g') s x := (is_bounded_bilinear_map_inner.has_fderiv_at (f x, g x)).comp_has_fderiv_within_at x (hf.prod hg) lemma has_strict_fderiv_at.inner (hf : has_strict_fderiv_at f f' x) (hg : has_strict_fderiv_at g g' x) : - has_strict_fderiv_at (λ t, ⟪f t, g t⟫) ((fderiv_inner_clm (f x, g x)).comp $ f'.prod g') x := + has_strict_fderiv_at (λ t, ⟪f t, g t⟫) ((fderiv_inner_clm 𝕜 (f x, g x)).comp $ f'.prod g') x := (is_bounded_bilinear_map_inner.has_strict_fderiv_at (f x, g x)).comp x (hf.prod hg) lemma has_fderiv_at.inner (hf : has_fderiv_at f f' x) (hg : has_fderiv_at g g' x) : - has_fderiv_at (λ t, ⟪f t, g t⟫) ((fderiv_inner_clm (f x, g x)).comp $ f'.prod g') x := + has_fderiv_at (λ t, ⟪f t, g t⟫) ((fderiv_inner_clm 𝕜 (f x, g x)).comp $ f'.prod g') x := (is_bounded_bilinear_map_inner.has_fderiv_at (f x, g x)).comp x (hf.prod hg) lemma has_deriv_within_at.inner {f g : ℝ → E} {f' g' : E} {s : set ℝ} {x : ℝ} (hf : has_deriv_within_at f f' s x) (hg : has_deriv_within_at g g' s x) : has_deriv_within_at (λ t, ⟪f t, g t⟫) (⟪f x, g'⟫ + ⟪f', g x⟫) s x := -by simpa using (hf.has_fderiv_within_at.inner hg.has_fderiv_within_at).has_deriv_within_at +by simpa using (hf.has_fderiv_within_at.inner 𝕜 hg.has_fderiv_within_at).has_deriv_within_at lemma has_deriv_at.inner {f g : ℝ → E} {f' g' : E} {x : ℝ} : has_deriv_at f f' x → has_deriv_at g g' x → has_deriv_at (λ t, ⟪f t, g t⟫) (⟪f x, g'⟫ + ⟪f', g x⟫) x := -by simpa only [← has_deriv_within_at_univ] using has_deriv_within_at.inner +by simpa only [← has_deriv_within_at_univ] using has_deriv_within_at.inner 𝕜 lemma differentiable_within_at.inner (hf : differentiable_within_at ℝ f s x) (hg : differentiable_within_at ℝ g s x) : @@ -111,90 +112,90 @@ lemma differentiable_at.inner (hf : differentiable_at ℝ f x) (hg : differentia lemma differentiable_on.inner (hf : differentiable_on ℝ f s) (hg : differentiable_on ℝ g s) : differentiable_on ℝ (λ x, ⟪f x, g x⟫) s := -λ x hx, (hf x hx).inner (hg x hx) +λ x hx, (hf x hx).inner 𝕜 (hg x hx) lemma differentiable.inner (hf : differentiable ℝ f) (hg : differentiable ℝ g) : differentiable ℝ (λ x, ⟪f x, g x⟫) := -λ x, (hf x).inner (hg x) +λ x, (hf x).inner 𝕜 (hg x) lemma fderiv_inner_apply (hf : differentiable_at ℝ f x) (hg : differentiable_at ℝ g x) (y : G) : fderiv ℝ (λ t, ⟪f t, g t⟫) x y = ⟪f x, fderiv ℝ g x y⟫ + ⟪fderiv ℝ f x y, g x⟫ := -by { rw [(hf.has_fderiv_at.inner hg.has_fderiv_at).fderiv], refl } +by { rw [(hf.has_fderiv_at.inner 𝕜 hg.has_fderiv_at).fderiv], refl } lemma deriv_inner_apply {f g : ℝ → E} {x : ℝ} (hf : differentiable_at ℝ f x) (hg : differentiable_at ℝ g x) : deriv (λ t, ⟪f t, g t⟫) x = ⟪f x, deriv g x⟫ + ⟪deriv f x, g x⟫ := -(hf.has_deriv_at.inner hg.has_deriv_at).deriv +(hf.has_deriv_at.inner 𝕜 hg.has_deriv_at).deriv lemma cont_diff_norm_sq : cont_diff ℝ n (λ x : E, ‖x‖ ^ 2) := begin - simp only [sq, ← inner_self_eq_norm_mul_norm], - exact (re_clm : 𝕜 →L[ℝ] ℝ).cont_diff.comp (cont_diff_id.inner cont_diff_id) + simp only [sq, ← @inner_self_eq_norm_mul_norm 𝕜], + exact (re_clm : 𝕜 →L[ℝ] ℝ).cont_diff.comp (cont_diff_id.inner 𝕜 cont_diff_id) end lemma cont_diff.norm_sq (hf : cont_diff ℝ n f) : cont_diff ℝ n (λ x, ‖f x‖ ^ 2) := -cont_diff_norm_sq.comp hf +(cont_diff_norm_sq 𝕜).comp hf lemma cont_diff_within_at.norm_sq (hf : cont_diff_within_at ℝ n f s x) : cont_diff_within_at ℝ n (λ y, ‖f y‖ ^ 2) s x := -cont_diff_norm_sq.cont_diff_at.comp_cont_diff_within_at x hf +(cont_diff_norm_sq 𝕜).cont_diff_at.comp_cont_diff_within_at x hf lemma cont_diff_at.norm_sq (hf : cont_diff_at ℝ n f x) : cont_diff_at ℝ n (λ y, ‖f y‖ ^ 2) x := -hf.norm_sq +hf.norm_sq 𝕜 lemma cont_diff_at_norm {x : E} (hx : x ≠ 0) : cont_diff_at ℝ n norm x := have ‖id x‖ ^ 2 ≠ 0, from pow_ne_zero _ (norm_pos_iff.2 hx).ne', -by simpa only [id, sqrt_sq, norm_nonneg] using cont_diff_at_id.norm_sq.sqrt this +by simpa only [id, sqrt_sq, norm_nonneg] using (cont_diff_at_id.norm_sq 𝕜).sqrt this lemma cont_diff_at.norm (hf : cont_diff_at ℝ n f x) (h0 : f x ≠ 0) : cont_diff_at ℝ n (λ y, ‖f y‖) x := -(cont_diff_at_norm h0).comp x hf +(cont_diff_at_norm 𝕜 h0).comp x hf lemma cont_diff_at.dist (hf : cont_diff_at ℝ n f x) (hg : cont_diff_at ℝ n g x) (hne : f x ≠ g x) : cont_diff_at ℝ n (λ y, dist (f y) (g y)) x := -by { simp only [dist_eq_norm], exact (hf.sub hg).norm (sub_ne_zero.2 hne) } +by { simp only [dist_eq_norm], exact (hf.sub hg).norm 𝕜 (sub_ne_zero.2 hne) } lemma cont_diff_within_at.norm (hf : cont_diff_within_at ℝ n f s x) (h0 : f x ≠ 0) : cont_diff_within_at ℝ n (λ y, ‖f y‖) s x := -(cont_diff_at_norm h0).comp_cont_diff_within_at x hf +(cont_diff_at_norm 𝕜 h0).comp_cont_diff_within_at x hf lemma cont_diff_within_at.dist (hf : cont_diff_within_at ℝ n f s x) (hg : cont_diff_within_at ℝ n g s x) (hne : f x ≠ g x) : cont_diff_within_at ℝ n (λ y, dist (f y) (g y)) s x := -by { simp only [dist_eq_norm], exact (hf.sub hg).norm (sub_ne_zero.2 hne) } +by { simp only [dist_eq_norm], exact (hf.sub hg).norm 𝕜 (sub_ne_zero.2 hne) } lemma cont_diff_on.norm_sq (hf : cont_diff_on ℝ n f s) : cont_diff_on ℝ n (λ y, ‖f y‖ ^ 2) s := -(λ x hx, (hf x hx).norm_sq) +(λ x hx, (hf x hx).norm_sq 𝕜) lemma cont_diff_on.norm (hf : cont_diff_on ℝ n f s) (h0 : ∀ x ∈ s, f x ≠ 0) : cont_diff_on ℝ n (λ y, ‖f y‖) s := -λ x hx, (hf x hx).norm (h0 x hx) +λ x hx, (hf x hx).norm 𝕜 (h0 x hx) lemma cont_diff_on.dist (hf : cont_diff_on ℝ n f s) (hg : cont_diff_on ℝ n g s) (hne : ∀ x ∈ s, f x ≠ g x) : cont_diff_on ℝ n (λ y, dist (f y) (g y)) s := -λ x hx, (hf x hx).dist (hg x hx) (hne x hx) +λ x hx, (hf x hx).dist 𝕜 (hg x hx) (hne x hx) lemma cont_diff.norm (hf : cont_diff ℝ n f) (h0 : ∀ x, f x ≠ 0) : cont_diff ℝ n (λ y, ‖f y‖) := -cont_diff_iff_cont_diff_at.2 $ λ x, hf.cont_diff_at.norm (h0 x) +cont_diff_iff_cont_diff_at.2 $ λ x, hf.cont_diff_at.norm 𝕜 (h0 x) lemma cont_diff.dist (hf : cont_diff ℝ n f) (hg : cont_diff ℝ n g) (hne : ∀ x, f x ≠ g x) : cont_diff ℝ n (λ y, dist (f y) (g y)) := cont_diff_iff_cont_diff_at.2 $ - λ x, hf.cont_diff_at.dist hg.cont_diff_at (hne x) + λ x, hf.cont_diff_at.dist 𝕜 hg.cont_diff_at (hne x) omit 𝕜 lemma has_strict_fderiv_at_norm_sq (x : F) : has_strict_fderiv_at (λ x, ‖x‖ ^ 2) (bit0 (innerSL ℝ x)) x := begin - simp only [sq, ← inner_self_eq_norm_mul_norm], - convert (has_strict_fderiv_at_id x).inner (has_strict_fderiv_at_id x), + simp only [sq, ← @inner_self_eq_norm_mul_norm ℝ], + convert (has_strict_fderiv_at_id x).inner ℝ (has_strict_fderiv_at_id x), ext y, simp [bit0, real_inner_comm], end @@ -202,54 +203,54 @@ include 𝕜 lemma differentiable_at.norm_sq (hf : differentiable_at ℝ f x) : differentiable_at ℝ (λ y, ‖f y‖ ^ 2) x := -(cont_diff_at_id.norm_sq.differentiable_at le_rfl).comp x hf +((cont_diff_at_id.norm_sq 𝕜).differentiable_at le_rfl).comp x hf lemma differentiable_at.norm (hf : differentiable_at ℝ f x) (h0 : f x ≠ 0) : differentiable_at ℝ (λ y, ‖f y‖) x := -((cont_diff_at_norm h0).differentiable_at le_rfl).comp x hf +((cont_diff_at_norm 𝕜 h0).differentiable_at le_rfl).comp x hf lemma differentiable_at.dist (hf : differentiable_at ℝ f x) (hg : differentiable_at ℝ g x) (hne : f x ≠ g x) : differentiable_at ℝ (λ y, dist (f y) (g y)) x := -by { simp only [dist_eq_norm], exact (hf.sub hg).norm (sub_ne_zero.2 hne) } +by { simp only [dist_eq_norm], exact (hf.sub hg).norm 𝕜 (sub_ne_zero.2 hne) } lemma differentiable.norm_sq (hf : differentiable ℝ f) : differentiable ℝ (λ y, ‖f y‖ ^ 2) := -λ x, (hf x).norm_sq +λ x, (hf x).norm_sq 𝕜 lemma differentiable.norm (hf : differentiable ℝ f) (h0 : ∀ x, f x ≠ 0) : differentiable ℝ (λ y, ‖f y‖) := -λ x, (hf x).norm (h0 x) +λ x, (hf x).norm 𝕜 (h0 x) lemma differentiable.dist (hf : differentiable ℝ f) (hg : differentiable ℝ g) (hne : ∀ x, f x ≠ g x) : differentiable ℝ (λ y, dist (f y) (g y)) := -λ x, (hf x).dist (hg x) (hne x) +λ x, (hf x).dist 𝕜 (hg x) (hne x) lemma differentiable_within_at.norm_sq (hf : differentiable_within_at ℝ f s x) : differentiable_within_at ℝ (λ y, ‖f y‖ ^ 2) s x := -(cont_diff_at_id.norm_sq.differentiable_at le_rfl).comp_differentiable_within_at x hf +((cont_diff_at_id.norm_sq 𝕜).differentiable_at le_rfl).comp_differentiable_within_at x hf lemma differentiable_within_at.norm (hf : differentiable_within_at ℝ f s x) (h0 : f x ≠ 0) : differentiable_within_at ℝ (λ y, ‖f y‖) s x := -((cont_diff_at_id.norm h0).differentiable_at le_rfl).comp_differentiable_within_at x hf +((cont_diff_at_id.norm 𝕜 h0).differentiable_at le_rfl).comp_differentiable_within_at x hf lemma differentiable_within_at.dist (hf : differentiable_within_at ℝ f s x) (hg : differentiable_within_at ℝ g s x) (hne : f x ≠ g x) : differentiable_within_at ℝ (λ y, dist (f y) (g y)) s x := -by { simp only [dist_eq_norm], exact (hf.sub hg).norm (sub_ne_zero.2 hne) } +by { simp only [dist_eq_norm], exact (hf.sub hg).norm 𝕜 (sub_ne_zero.2 hne) } lemma differentiable_on.norm_sq (hf : differentiable_on ℝ f s) : differentiable_on ℝ (λ y, ‖f y‖ ^ 2) s := -λ x hx, (hf x hx).norm_sq +λ x hx, (hf x hx).norm_sq 𝕜 lemma differentiable_on.norm (hf : differentiable_on ℝ f s) (h0 : ∀ x ∈ s, f x ≠ 0) : differentiable_on ℝ (λ y, ‖f y‖) s := -λ x hx, (hf x hx).norm (h0 x hx) +λ x hx, (hf x hx).norm 𝕜 (h0 x hx) lemma differentiable_on.dist (hf : differentiable_on ℝ f s) (hg : differentiable_on ℝ g s) (hne : ∀ x ∈ s, f x ≠ g x) : differentiable_on ℝ (λ y, dist (f y) (g y)) s := -λ x hx, (hf x hx).dist (hg x hx) (hne x hx) +λ x hx, (hf x hx).dist 𝕜 (hg x hx) (hne x hx) end deriv_inner @@ -338,7 +339,7 @@ section diffeomorph_unit_ball open metric (hiding mem_nhds_iff) -variables {n : ℕ∞} {E : Type*} [inner_product_space ℝ E] +variables {n : ℕ∞} {E : Type*} [normed_add_comm_group E] [inner_product_space ℝ E] lemma cont_diff_homeomorph_unit_ball : cont_diff ℝ n $ λ (x : E), (homeomorph_unit_ball x : E) := @@ -346,7 +347,7 @@ begin suffices : cont_diff ℝ n (λ x, (1 + ‖x‖^2).sqrt⁻¹), { exact this.smul cont_diff_id, }, have h : ∀ (x : E), 0 < 1 + ‖x‖ ^ 2 := λ x, by positivity, refine cont_diff.inv _ (λ x, real.sqrt_ne_zero'.mpr (h x)), - exact (cont_diff_const.add cont_diff_norm_sq).sqrt (λ x, (h x).ne.symm), + exact (cont_diff_const.add $ cont_diff_norm_sq ℝ).sqrt (λ x, (h x).ne.symm), end lemma cont_diff_on_homeomorph_unit_ball_symm @@ -367,7 +368,7 @@ begin ← sq_lt_sq, one_pow, ← sub_pos] at hy, refine cont_diff_at.inv _ (real.sqrt_ne_zero'.mpr h), refine cont_diff_at.comp _ (cont_diff_at_sqrt h.ne.symm) _, - exact cont_diff_at_const.sub cont_diff_norm_sq.cont_diff_at, + exact cont_diff_at_const.sub (cont_diff_norm_sq ℝ).cont_diff_at, end end diffeomorph_unit_ball diff --git a/src/analysis/inner_product_space/conformal_linear_map.lean b/src/analysis/inner_product_space/conformal_linear_map.lean index d02ee93c30e53..e252ae03ccca7 100644 --- a/src/analysis/inner_product_space/conformal_linear_map.lean +++ b/src/analysis/inner_product_space/conformal_linear_map.lean @@ -12,7 +12,9 @@ import analysis.inner_product_space.basic In an inner product space, a map is conformal iff it preserves inner products up to a scalar factor. -/ -variables {E F : Type*} [inner_product_space ℝ E] [inner_product_space ℝ F] +variables {E F : Type*} +variables [normed_add_comm_group E] [normed_add_comm_group F] +variables [inner_product_space ℝ E] [inner_product_space ℝ F] open linear_isometry continuous_linear_map open_locale real_inner_product_space diff --git a/src/analysis/inner_product_space/dual.lean b/src/analysis/inner_product_space/dual.lean index 3edc5c7a7c465..adbbc2ad54936 100644 --- a/src/analysis/inner_product_space/dual.lean +++ b/src/analysis/inner_product_space/dual.lean @@ -42,7 +42,7 @@ namespace inner_product_space open is_R_or_C continuous_linear_map variables (𝕜 : Type*) -variables (E : Type*) [is_R_or_C 𝕜] [inner_product_space 𝕜 E] +variables (E : Type*) [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] local notation `⟪`x`, `y`⟫` := @inner 𝕜 E _ x y local postfix `†`:90 := star_ring_end _ diff --git a/src/analysis/inner_product_space/gram_schmidt_ortho.lean b/src/analysis/inner_product_space/gram_schmidt_ortho.lean index fbbea0aac7209..da9d09de3772d 100644 --- a/src/analysis/inner_product_space/gram_schmidt_ortho.lean +++ b/src/analysis/inner_product_space/gram_schmidt_ortho.lean @@ -38,7 +38,7 @@ and outputs a set of orthogonal vectors which have the same span. open_locale big_operators open finset submodule finite_dimensional -variables (𝕜 : Type*) {E : Type*} [is_R_or_C 𝕜] [inner_product_space 𝕜 E] +variables (𝕜 : Type*) {E : Type*} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] variables {ι : Type*} [linear_order ι] [locally_finite_order_bot ι] [is_well_order ι (<)] local attribute [instance] is_well_order.to_has_well_founded diff --git a/src/analysis/inner_product_space/l2_space.lean b/src/analysis/inner_product_space/l2_space.lean index 825030c977a92..1dba544d0c830 100644 --- a/src/analysis/inner_product_space/l2_space.lean +++ b/src/analysis/inner_product_space/l2_space.lean @@ -85,8 +85,9 @@ open_locale big_operators nnreal ennreal classical complex_conjugate topology noncomputable theory variables {ι : Type*} -variables {𝕜 : Type*} [is_R_or_C 𝕜] {E : Type*} [inner_product_space 𝕜 E] [cplt : complete_space E] -variables {G : ι → Type*} [Π i, inner_product_space 𝕜 (G i)] +variables {𝕜 : Type*} [is_R_or_C 𝕜] {E : Type*} +variables [normed_add_comm_group E] [inner_product_space 𝕜 E] [cplt : complete_space E] +variables {G : ι → Type*} [Π i, normed_add_comm_group (G i)] [Π i, inner_product_space 𝕜 (G i)] local notation `⟪`x`, `y`⟫` := @inner 𝕜 _ _ x y notation `ℓ²(`ι`, `𝕜`)` := lp (λ i : ι, 𝕜) 2 @@ -111,7 +112,7 @@ instance : inner_product_space 𝕜 (lp G 2) := calc ‖f‖ ^ 2 = ‖f‖ ^ (2:ℝ≥0∞).to_real : by norm_cast ... = ∑' i, ‖f i‖ ^ (2:ℝ≥0∞).to_real : lp.norm_rpow_eq_tsum _ f ... = ∑' i, ‖f i‖ ^ 2 : by norm_cast - ... = ∑' i, re ⟪f i, f i⟫ : by simp only [norm_sq_eq_inner] + ... = ∑' i, re ⟪f i, f i⟫ : by simp only [@norm_sq_eq_inner 𝕜] ... = re (∑' i, ⟪f i, f i⟫) : (is_R_or_C.re_clm.map_tsum _).symm ... = _ : by congr, { norm_num }, @@ -160,7 +161,7 @@ begin end lemma inner_single_right (i : ι) (a : G i) (f : lp G 2) : ⟪f, lp.single 2 i a⟫ = ⟪f i, a⟫ := -by simpa [inner_conj_symm] using congr_arg conj (inner_single_left i a f) +by simpa [inner_conj_symm] using congr_arg conj (@inner_single_left _ 𝕜 _ _ _ _ i a f) end lp @@ -425,6 +426,7 @@ begin exact (↑(b.repr.symm.to_continuous_linear_equiv) : ℓ²(ι, 𝕜) →L[𝕜] E).has_sum this }, ext i, apply b.repr.injective, + letI : normed_space 𝕜 ↥(lp (λ i : ι, 𝕜) 2) := by apply_instance, have : lp.single 2 i (f i * 1) = f i • lp.single 2 i 1 := lp.single_smul 2 i (1:𝕜) (f i), rw mul_one at this, rw [linear_isometry_equiv.map_smul, b.repr_self, ← this, diff --git a/src/analysis/inner_product_space/lax_milgram.lean b/src/analysis/inner_product_space/lax_milgram.lean index 24753555f05ef..c6f4f84191f2a 100644 --- a/src/analysis/inner_product_space/lax_milgram.lean +++ b/src/analysis/inner_product_space/lax_milgram.lean @@ -40,10 +40,10 @@ open_locale real_inner_product_space nnreal universe u namespace is_coercive -variables {V : Type u} [inner_product_space ℝ V] [complete_space V] +variables {V : Type u} [normed_add_comm_group V] [inner_product_space ℝ V] [complete_space V] variables {B : V →L[ℝ] V →L[ℝ] ℝ} -local postfix `♯`:1025 := @continuous_linear_map_of_bilin ℝ V _ _ _ +local postfix `♯`:1025 := @continuous_linear_map_of_bilin ℝ V _ _ _ _ lemma bounded_below (coercive : is_coercive B) : ∃ C, 0 < C ∧ ∀ v, C * ‖v‖ ≤ ‖B♯ v‖ := diff --git a/src/analysis/inner_product_space/orientation.lean b/src/analysis/inner_product_space/orientation.lean index 17a6ec884d5b9..d962ce8bfe730 100644 --- a/src/analysis/inner_product_space/orientation.lean +++ b/src/analysis/inner_product_space/orientation.lean @@ -35,7 +35,7 @@ This file provides definitions and proves lemmas about orientations of real inne noncomputable theory -variables {E : Type*} [inner_product_space ℝ E] +variables {E : Type*} [normed_add_comm_group E] [inner_product_space ℝ E] open finite_dimensional open_locale big_operators real_inner_product_space @@ -317,7 +317,8 @@ lemma abs_volume_form_apply_of_orthonormal (v : orthonormal_basis (fin n) ℝ E) |o.volume_form v| = 1 := by simpa [o.volume_form_robust' v v] using congr_arg abs v.to_basis.det_self -lemma volume_form_map {F : Type*} [inner_product_space ℝ F] [fact (finrank ℝ F = n)] +lemma volume_form_map {F : Type*} + [normed_add_comm_group F] [inner_product_space ℝ F] [fact (finrank ℝ F = n)] (φ : E ≃ₗᵢ[ℝ] F) (x : fin n → F) : (orientation.map (fin n) φ.to_linear_equiv o).volume_form x = o.volume_form (φ.symm ∘ x) := begin diff --git a/src/analysis/inner_product_space/pi_L2.lean b/src/analysis/inner_product_space/pi_L2.lean index deec2a2192836..4c865e310c985 100644 --- a/src/analysis/inner_product_space/pi_L2.lean +++ b/src/analysis/inner_product_space/pi_L2.lean @@ -57,10 +57,11 @@ open_locale big_operators uniformity topology nnreal ennreal complex_conjugate d noncomputable theory variables {ι : Type*} {ι' : Type*} -variables {𝕜 : Type*} [is_R_or_C 𝕜] {E : Type*} [inner_product_space 𝕜 E] -variables {E' : Type*} [inner_product_space 𝕜 E'] -variables {F : Type*} [inner_product_space ℝ F] -variables {F' : Type*} [inner_product_space ℝ F'] +variables {𝕜 : Type*} [is_R_or_C 𝕜] +variables {E : Type*} [normed_add_comm_group E] [inner_product_space 𝕜 E] +variables {E' : Type*} [normed_add_comm_group E'] [inner_product_space 𝕜 E'] +variables {F : Type*} [normed_add_comm_group F] [inner_product_space ℝ F] +variables {F' : Type*} [normed_add_comm_group F'] [inner_product_space ℝ F'] local notation `⟪`x`, `y`⟫` := @inner 𝕜 _ _ x y /- @@ -69,9 +70,9 @@ then `Π i, f i` is an inner product space as well. Since `Π i, f i` is endowed we use instead `pi_Lp 2 f` for the product space, which is endowed with the `L^2` norm. -/ instance pi_Lp.inner_product_space {ι : Type*} [fintype ι] (f : ι → Type*) - [Π i, inner_product_space 𝕜 (f i)] : inner_product_space 𝕜 (pi_Lp 2 f) := -{ to_normed_add_comm_group := infer_instance, - inner := λ x y, ∑ i, inner (x i) (y i), + [Π i, normed_add_comm_group (f i)] [Π i, inner_product_space 𝕜 (f i)] : + inner_product_space 𝕜 (pi_Lp 2 f) := +{ inner := λ x y, ∑ i, inner (x i) (y i), norm_sq_eq_inner := λ x, by simp only [pi_Lp.norm_sq_eq_of_L2, add_monoid_hom.map_sum, ← norm_sq_eq_inner, one_div], conj_symm := @@ -91,7 +92,7 @@ instance pi_Lp.inner_product_space {ι : Type*} [fintype ι] (f : ι → Type*) by simp only [finset.mul_sum, inner_smul_left] } @[simp] lemma pi_Lp.inner_apply {ι : Type*} [fintype ι] {f : ι → Type*} - [Π i, inner_product_space 𝕜 (f i)] (x y : pi_Lp 2 f) : + [Π i, normed_add_comm_group (f i)] [Π i, inner_product_space 𝕜 (f i)] (x y : pi_Lp 2 f) : ⟪x, y⟫ = ∑ i, ⟪x i, y i⟫ := rfl @@ -150,7 +151,7 @@ def direct_sum.is_internal.isometry_L2_of_orthogonal_family begin let e₁ := direct_sum.linear_equiv_fun_on_fintype 𝕜 ι (λ i, V i), let e₂ := linear_equiv.of_bijective (direct_sum.coe_linear_map V) hV, - refine (e₂.symm.trans e₁).isometry_of_inner _, + refine linear_equiv.isometry_of_inner (e₂.symm.trans e₁) _, suffices : ∀ v w, ⟪v, w⟫ = ⟪e₂ (e₁.symm v), e₂ (e₁.symm w)⟫, { intros v₀ w₀, convert this (e₁ (e₂.symm v₀)) (e₁ (e₂.symm w₀)); @@ -344,16 +345,17 @@ by simpa only [b.repr_apply_apply, inner_orthogonal_projection_eq_of_mem_left] using (b.sum_repr (orthogonal_projection U x)).symm /-- Mapping an orthonormal basis along a `linear_isometry_equiv`. -/ -protected def map {G : Type*} [inner_product_space 𝕜 G] (b : orthonormal_basis ι 𝕜 E) +protected def map {G : Type*} + [normed_add_comm_group G] [inner_product_space 𝕜 G] (b : orthonormal_basis ι 𝕜 E) (L : E ≃ₗᵢ[𝕜] G) : orthonormal_basis ι 𝕜 G := { repr := L.symm.trans b.repr } -@[simp] protected lemma map_apply {G : Type*} [inner_product_space 𝕜 G] +@[simp] protected lemma map_apply {G : Type*} [normed_add_comm_group G] [inner_product_space 𝕜 G] (b : orthonormal_basis ι 𝕜 E) (L : E ≃ₗᵢ[𝕜] G) (i : ι) : b.map L i = L (b i) := rfl -@[simp] protected lemma to_basis_map {G : Type*} [inner_product_space 𝕜 G] +@[simp] protected lemma to_basis_map {G : Type*} [normed_add_comm_group G] [inner_product_space 𝕜 G] (b : orthonormal_basis ι 𝕜 E) (L : E ≃ₗᵢ[𝕜] G) : (b.map L).to_basis = b.to_basis.map L.to_linear_equiv := rfl @@ -737,7 +739,7 @@ def orthonormal_basis.from_orthogonal_span_singleton section linear_isometry -variables {V : Type*} [inner_product_space 𝕜 V] [finite_dimensional 𝕜 V] +variables {V : Type*} [normed_add_comm_group V] [inner_product_space 𝕜 V] [finite_dimensional 𝕜 V] variables {S : submodule 𝕜 V} {L : S →ₗᵢ[𝕜] V} diff --git a/src/analysis/inner_product_space/positive.lean b/src/analysis/inner_product_space/positive.lean index 46fd7ddd5b7ad..3b7e86b17d114 100644 --- a/src/analysis/inner_product_space/positive.lean +++ b/src/analysis/inner_product_space/positive.lean @@ -38,8 +38,10 @@ open_locale inner_product complex_conjugate namespace continuous_linear_map -variables {𝕜 E F : Type*} [is_R_or_C 𝕜] [inner_product_space 𝕜 E] [inner_product_space 𝕜 F] - [complete_space E] [complete_space F] +variables {𝕜 E F : Type*} [is_R_or_C 𝕜] +variables [normed_add_comm_group E] [normed_add_comm_group F] +variables [inner_product_space 𝕜 E] [inner_product_space 𝕜 F] +variables [complete_space E] [complete_space F] local notation `⟪`x`, `y`⟫` := @inner 𝕜 _ _ x y /-- A continuous linear endomorphism `T` of a Hilbert space is **positive** if it is self adjoint @@ -111,7 +113,7 @@ end section complex -variables {E' : Type*} [inner_product_space ℂ E'] [complete_space E'] +variables {E' : Type*} [normed_add_comm_group E'] [inner_product_space ℂ E'] [complete_space E'] lemma is_positive_iff_complex (T : E' →L[ℂ] E') : is_positive T ↔ ∀ x, (re ⟪T x, x⟫_ℂ : ℂ) = ⟪T x, x⟫_ℂ ∧ 0 ≤ re ⟪T x, x⟫_ℂ := diff --git a/src/analysis/inner_product_space/projection.lean b/src/analysis/inner_product_space/projection.lean index 2b1852d5fbbac..e2e124bec373a 100644 --- a/src/analysis/inner_product_space/projection.lean +++ b/src/analysis/inner_product_space/projection.lean @@ -43,6 +43,7 @@ open is_R_or_C real filter linear_map (ker range) open_locale big_operators topology variables {𝕜 E F : Type*} [is_R_or_C 𝕜] +variables [normed_add_comm_group E] [normed_add_comm_group F] variables [inner_product_space 𝕜 E] [inner_product_space ℝ F] local notation `⟪`x`, `y`⟫` := @inner 𝕜 _ _ x y local notation `absR` := has_abs.abs @@ -116,7 +117,7 @@ begin have eq₂ : u + u - (wq + wp) = a + b, show u + u - (wq + wp) = (u - wq) + (u - wp), abel, rw [eq₁, eq₂], end - ... = 2 * (‖a‖ * ‖a‖ + ‖b‖ * ‖b‖) : parallelogram_law_with_norm _ _, + ... = 2 * (‖a‖ * ‖a‖ + ‖b‖ * ‖b‖) : parallelogram_law_with_norm ℝ _ _, have eq : δ ≤ ‖u - half • (wq + wp)‖, { rw smul_add, apply δ_le', apply h₂, @@ -204,7 +205,7 @@ begin end ... = ‖u - v‖^2 - 2 * θ * inner (u - v) (w - v) + θ*θ*‖w - v‖^2 : begin - rw [norm_sub_sq, inner_smul_right, norm_smul], + rw [@norm_sub_sq ℝ, inner_smul_right, norm_smul], simp only [sq], show ‖u-v‖*‖u-v‖-2*(θ*inner(u-v)(w-v))+absR (θ)*‖w-v‖*(absR (θ)*‖w-v‖)= ‖u-v‖*‖u-v‖-2*θ*inner(u-v)(w-v)+θ*θ*(‖w-v‖*‖w-v‖), @@ -246,7 +247,7 @@ begin ‖u - v‖ * ‖u - v‖ ≤ ‖u - v‖ * ‖u - v‖ - 2 * inner (u - v) ((w:F) - v) : by linarith ... ≤ ‖u - v‖^2 - 2 * inner (u - v) ((w:F) - v) + ‖(w:F) - v‖^2 : by { rw sq, refine le_add_of_nonneg_right _, exact sq_nonneg _ } - ... = ‖(u - v) - (w - v)‖^2 : (norm_sub_sq _ _).symm + ... = ‖(u - v) - (w - v)‖^2 : (@norm_sub_sq ℝ _ _ _ _ _ _).symm ... = ‖u - w‖ * ‖u - w‖ : by { have : (u - v) - (w - v) = u - w, abel, rw [this, sq] } }, { show (⨅ (w : K), ‖u - w‖) ≤ (λw:K, ‖u - w‖) ⟨v, hv⟩, @@ -382,7 +383,7 @@ lemma eq_orthogonal_projection_fn_of_mem_of_inner_eq_zero {u v : E} (hvm : v ∈ K) (hvo : ∀ w ∈ K, ⟪u - v, w⟫ = 0) : orthogonal_projection_fn K u = v := begin - rw [←sub_eq_zero, ←inner_self_eq_zero], + rw [←sub_eq_zero, ←@inner_self_eq_zero 𝕜], have hvs : orthogonal_projection_fn K u - v ∈ K := submodule.sub_mem K (orthogonal_projection_fn_mem u) hvm, have huo : ⟪u - orthogonal_projection_fn K u, orthogonal_projection_fn K u - v⟫ = 0 := @@ -499,8 +500,10 @@ begin { simp } end -lemma linear_isometry.map_orthogonal_projection {E E' : Type*} [inner_product_space 𝕜 E] - [inner_product_space 𝕜 E'] (f : E →ₗᵢ[𝕜] E') (p : submodule 𝕜 E) [complete_space p] +lemma linear_isometry.map_orthogonal_projection {E E' : Type*} + [normed_add_comm_group E] [normed_add_comm_group E'] + [inner_product_space 𝕜 E] [inner_product_space 𝕜 E'] + (f : E →ₗᵢ[𝕜] E') (p : submodule 𝕜 E) [complete_space p] (x : E) : f (orthogonal_projection p x) = orthogonal_projection (p.map f.to_linear_map) (f x) := begin @@ -511,8 +514,10 @@ begin rw [← f.map_sub, f.inner_map_map, orthogonal_projection_inner_eq_zero x x' hx'] end -lemma linear_isometry.map_orthogonal_projection' {E E' : Type*} [inner_product_space 𝕜 E] - [inner_product_space 𝕜 E'] (f : E →ₗᵢ[𝕜] E') (p : submodule 𝕜 E) [complete_space p] +lemma linear_isometry.map_orthogonal_projection' {E E' : Type*} + [normed_add_comm_group E] [normed_add_comm_group E'] + [inner_product_space 𝕜 E] [inner_product_space 𝕜 E'] + (f : E →ₗᵢ[𝕜] E') (p : submodule 𝕜 E) [complete_space p] (x : E) : f (orthogonal_projection p x) = orthogonal_projection (p.map f) (f x) := begin @@ -524,8 +529,10 @@ begin end /-- Orthogonal projection onto the `submodule.map` of a subspace. -/ -lemma orthogonal_projection_map_apply {E E' : Type*} [inner_product_space 𝕜 E] - [inner_product_space 𝕜 E'] (f : E ≃ₗᵢ[𝕜] E') (p : submodule 𝕜 E) [complete_space p] +lemma orthogonal_projection_map_apply {E E' : Type*} + [normed_add_comm_group E] [normed_add_comm_group E'] + [inner_product_space 𝕜 E] [inner_product_space 𝕜 E'] + (f : E ≃ₗᵢ[𝕜] E') (p : submodule 𝕜 E) [complete_space p] (x : E') : (orthogonal_projection (p.map (f.to_linear_equiv : E →ₗ[𝕜] E')) x : E') = f (orthogonal_projection p (f.symm x)) := @@ -554,7 +561,7 @@ begin use ⟪v, w⟫ }, { intros x hx, obtain ⟨c, rfl⟩ := submodule.mem_span_singleton.mp hx, - have hv : ↑‖v‖ ^ 2 = ⟪v, v⟫ := by { norm_cast, simp [norm_sq_eq_inner] }, + have hv : ↑‖v‖ ^ 2 = ⟪v, v⟫ := by { norm_cast, simp [@norm_sq_eq_inner 𝕜] }, simp [inner_sub_left, inner_smul_left, inner_smul_right, map_div₀, mul_comm, hv, inner_product_space.conj_symm, hv] } end @@ -658,13 +665,17 @@ lemma reflection_mem_subspace_eq_self {x : E} (hx : x ∈ K) : reflection K x = (reflection_eq_self_iff x).mpr hx /-- Reflection in the `submodule.map` of a subspace. -/ -lemma reflection_map_apply {E E' : Type*} [inner_product_space 𝕜 E] [inner_product_space 𝕜 E'] +lemma reflection_map_apply {E E' : Type*} + [normed_add_comm_group E] [normed_add_comm_group E'] + [inner_product_space 𝕜 E] [inner_product_space 𝕜 E'] (f : E ≃ₗᵢ[𝕜] E') (K : submodule 𝕜 E) [complete_space K] (x : E') : reflection (K.map (f.to_linear_equiv : E →ₗ[𝕜] E')) x = f (reflection K (f.symm x)) := by simp [bit0, reflection_apply, orthogonal_projection_map_apply f K x] /-- Reflection in the `submodule.map` of a subspace. -/ -lemma reflection_map {E E' : Type*} [inner_product_space 𝕜 E] [inner_product_space 𝕜 E'] +lemma reflection_map {E E' : Type*} + [normed_add_comm_group E] [normed_add_comm_group E'] + [inner_product_space 𝕜 E] [inner_product_space 𝕜 E'] (f : E ≃ₗᵢ[𝕜] E') (K : submodule 𝕜 E) [complete_space K] : reflection (K.map (f.to_linear_equiv : E →ₗ[𝕜] E')) = f.symm.trans ((reflection K).trans f) := linear_isometry_equiv.ext $ reflection_map_apply f K diff --git a/src/analysis/inner_product_space/rayleigh.lean b/src/analysis/inner_product_space/rayleigh.lean index 9b36279f603f7..04807e8acdf04 100644 --- a/src/analysis/inner_product_space/rayleigh.lean +++ b/src/analysis/inner_product_space/rayleigh.lean @@ -35,7 +35,7 @@ A slightly more elaborate corollary is that if `E` is complete and `T` is a comp -/ variables {𝕜 : Type*} [is_R_or_C 𝕜] -variables {E : Type*} [inner_product_space 𝕜 E] +variables {E : Type*} [normed_add_comm_group E] [inner_product_space 𝕜 E] local notation `⟪`x`, `y`⟫` := @inner 𝕜 _ _ x y open_locale nnreal @@ -89,13 +89,13 @@ end continuous_linear_map namespace is_self_adjoint section real -variables {F : Type*} [inner_product_space ℝ F] +variables {F : Type*} [normed_add_comm_group F] [inner_product_space ℝ F] lemma _root_.linear_map.is_symmetric.has_strict_fderiv_at_re_apply_inner_self {T : F →L[ℝ] F} (hT : (T : F →ₗ[ℝ] F).is_symmetric) (x₀ : F) : has_strict_fderiv_at T.re_apply_inner_self (_root_.bit0 (innerSL ℝ (T x₀))) x₀ := begin - convert T.has_strict_fderiv_at.inner (has_strict_fderiv_at_id x₀), + convert T.has_strict_fderiv_at.inner _ (has_strict_fderiv_at_id x₀), ext y, simp_rw [_root_.bit0, continuous_linear_map.comp_apply, continuous_linear_map.add_apply, innerSL_apply, fderiv_inner_clm_apply, id.def, continuous_linear_map.prod_apply, diff --git a/src/analysis/inner_product_space/spectrum.lean b/src/analysis/inner_product_space/spectrum.lean index 631f947432a04..f312010c46b87 100644 --- a/src/analysis/inner_product_space/spectrum.lean +++ b/src/analysis/inner_product_space/spectrum.lean @@ -44,7 +44,7 @@ self-adjoint operator, spectral theorem, diagonalization theorem -/ variables {𝕜 : Type*} [is_R_or_C 𝕜] [dec_𝕜 : decidable_eq 𝕜] -variables {E : Type*} [inner_product_space 𝕜 E] +variables {E : Type*} [normed_add_comm_group E] [inner_product_space 𝕜 E] local notation `⟪`x`, `y`⟫` := @inner 𝕜 E _ x y diff --git a/src/analysis/inner_product_space/symmetric.lean b/src/analysis/inner_product_space/symmetric.lean index c78aff51c2f58..bd358edccdbbb 100644 --- a/src/analysis/inner_product_space/symmetric.lean +++ b/src/analysis/inner_product_space/symmetric.lean @@ -35,8 +35,10 @@ open is_R_or_C open_locale complex_conjugate variables {𝕜 E E' F G : Type*} [is_R_or_C 𝕜] -variables [inner_product_space 𝕜 E] [inner_product_space 𝕜 F] [inner_product_space 𝕜 G] -variables [inner_product_space ℝ E'] +variables [normed_add_comm_group E] [inner_product_space 𝕜 E] +variables [normed_add_comm_group F] [inner_product_space 𝕜 F] +variables [normed_add_comm_group G] [inner_product_space 𝕜 G] +variables [normed_add_comm_group E'] [inner_product_space ℝ E'] local notation `⟪`x`, `y`⟫` := @inner 𝕜 _ _ x y namespace linear_map @@ -89,7 +91,7 @@ lemma is_symmetric.continuous [complete_space E] {T : E →ₗ[𝕜] E} (hT : is begin -- We prove it by using the closed graph theorem refine T.continuous_of_seq_closed_graph (λ u x y hu hTu, _), - rw [←sub_eq_zero, ←inner_self_eq_zero], + rw [←sub_eq_zero, ←@inner_self_eq_zero 𝕜], have hlhs : ∀ k : ℕ, ⟪T (u k) - T x, y - T x⟫ = ⟪u k - x, T (y - T x)⟫ := by { intro k, rw [←T.map_sub, hT] }, refine tendsto_nhds_unique ((hTu.sub_const _).inner tendsto_const_nhds) _, @@ -119,7 +121,7 @@ lemma is_symmetric.restrict_invariant {T : E →ₗ[𝕜] E} (hT : is_symmetric λ v w, hT v w lemma is_symmetric.restrict_scalars {T : E →ₗ[𝕜] E} (hT : T.is_symmetric) : - @linear_map.is_symmetric ℝ E _ (inner_product_space.is_R_or_C_to_real 𝕜 E) + @linear_map.is_symmetric ℝ E _ _ (inner_product_space.is_R_or_C_to_real 𝕜 E) (@linear_map.restrict_scalars ℝ 𝕜 _ _ _ _ _ _ (inner_product_space.is_R_or_C_to_real 𝕜 E).to_module (inner_product_space.is_R_or_C_to_real 𝕜 E).to_module _ _ _ T) := @@ -128,7 +130,7 @@ lemma is_symmetric.restrict_scalars {T : E →ₗ[𝕜] E} (hT : T.is_symmetric) section complex variables {V : Type*} - [inner_product_space ℂ V] + [normed_add_comm_group V] [inner_product_space ℂ V] /-- A linear operator on a complex inner product space is symmetric precisely when `⟪T v, v⟫_ℂ` is real for all v.-/ diff --git a/src/analysis/inner_product_space/two_dim.lean b/src/analysis/inner_product_space/two_dim.lean index 2283a8740c7d4..a55a4456334a1 100644 --- a/src/analysis/inner_product_space/two_dim.lean +++ b/src/analysis/inner_product_space/two_dim.lean @@ -71,7 +71,7 @@ open finite_dimensional local attribute [instance] fact_finite_dimensional_of_finrank_eq_succ -variables {E : Type*} [inner_product_space ℝ E] [fact (finrank ℝ E = 2)] +variables {E : Type*} [normed_add_comm_group E] [inner_product_space ℝ E] [fact (finrank ℝ E = 2)] (o : orientation ℝ E (fin 2)) namespace orientation @@ -147,7 +147,8 @@ begin { simpa } end -lemma area_form_map {F : Type*} [inner_product_space ℝ F] [fact (finrank ℝ F = 2)] +lemma area_form_map {F : Type*} + [normed_add_comm_group F] [inner_product_space ℝ F] [fact (finrank ℝ F = 2)] (φ : E ≃ₗᵢ[ℝ] F) (x y : F) : (orientation.map (fin 2) φ.to_linear_equiv o).area_form x y = o.area_form (φ.symm x) (φ.symm y) := begin @@ -308,7 +309,8 @@ end (-o).right_angle_rotation = o.right_angle_rotation.trans (linear_isometry_equiv.neg ℝ) := linear_isometry_equiv.ext $ o.right_angle_rotation_neg_orientation -lemma right_angle_rotation_map {F : Type*} [inner_product_space ℝ F] [fact (finrank ℝ F = 2)] +lemma right_angle_rotation_map {F : Type*} + [normed_add_comm_group F] [inner_product_space ℝ F] [fact (finrank ℝ F = 2)] (φ : E ≃ₗᵢ[ℝ] F) (x : F) : (orientation.map (fin 2) φ.to_linear_equiv o).right_angle_rotation x = φ (o.right_angle_rotation (φ.symm x)) := @@ -335,7 +337,8 @@ begin rw [fact.out (finrank ℝ E = 2), fintype.card_fin] }, end -lemma right_angle_rotation_map' {F : Type*} [inner_product_space ℝ F] [fact (finrank ℝ F = 2)] +lemma right_angle_rotation_map' {F : Type*} + [normed_add_comm_group F] [inner_product_space ℝ F] [fact (finrank ℝ F = 2)] (φ : E ≃ₗᵢ[ℝ] F) : (orientation.map (fin 2) φ.to_linear_equiv o).right_angle_rotation = (φ.symm.trans o.right_angle_rotation).trans φ := @@ -556,7 +559,8 @@ begin simp, end -lemma kahler_map {F : Type*} [inner_product_space ℝ F] [fact (finrank ℝ F = 2)] +lemma kahler_map {F : Type*} + [normed_add_comm_group F] [inner_product_space ℝ F] [fact (finrank ℝ F = 2)] (φ : E ≃ₗᵢ[ℝ] F) (x y : F) : (orientation.map (fin 2) φ.to_linear_equiv o).kahler x y = o.kahler (φ.symm x) (φ.symm y) := by simp [kahler_apply_apply, area_form_map] diff --git a/src/analysis/matrix.lean b/src/analysis/matrix.lean index cfe6a970c3425..aef241badd380 100644 --- a/src/analysis/matrix.lean +++ b/src/analysis/matrix.lean @@ -453,7 +453,7 @@ begin rw [← nnreal.rpow_le_rpow_iff one_half_pos, ← nnreal.rpow_mul, mul_div_cancel' (1 : ℝ) two_ne_zero, nnreal.rpow_one, nnreal.mul_rpow], dsimp only, - have := @nnnorm_inner_le_nnnorm α _ _ _ + have := @nnnorm_inner_le_nnnorm α _ _ _ _ ((pi_Lp.equiv 2 (λ i, α)).symm (λ j, star (A i j))) ((pi_Lp.equiv 2 (λ i, α)).symm (λ k, B k j)), simpa only [pi_Lp.equiv_symm_apply, pi_Lp.inner_apply, diff --git a/src/analysis/quaternion.lean b/src/analysis/quaternion.lean index 762cb2d8533d1..1fedf745aa37b 100644 --- a/src/analysis/quaternion.lean +++ b/src/analysis/quaternion.lean @@ -41,8 +41,8 @@ lemma inner_self (a : ℍ) : ⟪a, a⟫ = norm_sq a := rfl lemma inner_def (a b : ℍ) : ⟪a, b⟫ = (a * b.conj).re := rfl -noncomputable instance : inner_product_space ℝ ℍ := -inner_product_space.of_core +noncomputable instance : normed_add_comm_group ℍ := +@inner_product_space.of_core.to_normed_add_comm_group ℝ ℍ _ _ _ { inner := has_inner.inner, conj_symm := λ x y, by simp [inner_def, mul_comm], nonneg_re := λ x, norm_sq_nonneg, @@ -50,6 +50,9 @@ inner_product_space.of_core add_left := λ x y z, by simp only [inner_def, add_mul, add_re], smul_left := λ x y r, by simp [inner_def] } +noncomputable instance : inner_product_space ℝ ℍ := +inner_product_space.of_core _ + lemma norm_sq_eq_norm_sq (a : ℍ) : norm_sq a = ‖a‖ * ‖a‖ := by rw [← inner_self, real_inner_self_eq_norm_mul_norm] diff --git a/src/analysis/von_neumann_algebra/basic.lean b/src/analysis/von_neumann_algebra/basic.lean index 365c0744733d1..84dc25e66e8be 100644 --- a/src/analysis/von_neumann_algebra/basic.lean +++ b/src/analysis/von_neumann_algebra/basic.lean @@ -65,7 +65,8 @@ Thus we can't say that the bounded operators `H →L[ℂ] H` form a `von_neumann and instead will use `⊤ : von_neumann_algebra H`. -/ @[nolint has_nonempty_instance] -structure von_neumann_algebra (H : Type u) [inner_product_space ℂ H] [complete_space H] extends +structure von_neumann_algebra (H : Type u) + [normed_add_comm_group H] [inner_product_space ℂ H] [complete_space H] extends star_subalgebra ℂ (H →L[ℂ] H) := (centralizer_centralizer' : set.centralizer (set.centralizer carrier) = carrier) @@ -78,7 +79,7 @@ or equivalently that it is closed in the weak and strong operator topologies.) add_decl_doc von_neumann_algebra.to_star_subalgebra namespace von_neumann_algebra -variables {H : Type u} [inner_product_space ℂ H] [complete_space H] +variables {H : Type u} [normed_add_comm_group H] [inner_product_space ℂ H] [complete_space H] instance : set_like (von_neumann_algebra H) (H →L[ℂ] H) := ⟨von_neumann_algebra.carrier, λ S T h, by cases S; cases T; congr'⟩ diff --git a/src/geometry/euclidean/angle/oriented/affine.lean b/src/geometry/euclidean/angle/oriented/affine.lean index ccb652414fe37..96caff2698604 100644 --- a/src/geometry/euclidean/angle/oriented/affine.lean +++ b/src/geometry/euclidean/angle/oriented/affine.lean @@ -26,8 +26,9 @@ open_locale affine euclidean_geometry real real_inner_product_space complex_conj namespace euclidean_geometry -variables {V : Type*} {P : Type*} [inner_product_space ℝ V] [metric_space P] -variables [normed_add_torsor V P] [hd2 : fact (finrank ℝ V = 2)] [module.oriented ℝ V (fin 2)] +variables {V : Type*} {P : Type*} + [normed_add_comm_group V] [inner_product_space ℝ V] [metric_space P] [normed_add_torsor V P] + [hd2 : fact (finrank ℝ V = 2)] [module.oriented ℝ V (fin 2)] include hd2 local notation `o` := module.oriented.positive_orientation @@ -596,10 +597,10 @@ begin { rw [@dist_eq_norm_vsub' V, @dist_eq_norm_vsub' V, ←mul_self_inj (norm_nonneg _) (norm_nonneg _), ←real_inner_self_eq_norm_mul_norm, ←real_inner_self_eq_norm_mul_norm] at hd, - simp_rw [vsub_midpoint, ←vsub_sub_vsub_cancel_left p₂ p₁ p, inner_sub_left, + simp_rw [vsub_midpoint, ←vsub_sub_vsub_cancel_left p₂ p₁ p, inner_sub_left, inner_add_right, inner_smul_right, hd, real_inner_comm (p -ᵥ p₁)], abel }, - rw [@orientation.inner_eq_zero_iff_eq_zero_or_eq_smul_rotation_pi_div_two V _ _ o, + rw [@orientation.inner_eq_zero_iff_eq_zero_or_eq_smul_rotation_pi_div_two V _ _ _ o, or_iff_right (vsub_ne_zero.2 h.symm)] at hi, rcases hi with ⟨r, hr⟩, rw [eq_comm, ←eq_vadd_iff_vsub_eq] at hr, diff --git a/src/geometry/euclidean/angle/oriented/basic.lean b/src/geometry/euclidean/angle/oriented/basic.lean index 4a79e60cf12f2..1a588447295da 100644 --- a/src/geometry/euclidean/angle/oriented/basic.lean +++ b/src/geometry/euclidean/angle/oriented/basic.lean @@ -39,7 +39,9 @@ namespace orientation local attribute [instance] fact_finite_dimensional_of_finrank_eq_succ local attribute [instance] complex.finrank_real_complex_fact -variables {V V' : Type*} [inner_product_space ℝ V] [inner_product_space ℝ V'] +variables {V V' : Type*} +variables [normed_add_comm_group V] [normed_add_comm_group V'] +variables [inner_product_space ℝ V] [inner_product_space ℝ V'] variables [fact (finrank ℝ V = 2)] [fact (finrank ℝ V' = 2)] (o : orientation ℝ V (fin 2)) local notation `ω` := o.area_form diff --git a/src/geometry/euclidean/angle/oriented/right_angle.lean b/src/geometry/euclidean/angle/oriented/right_angle.lean index 52462e99f46d8..a245fc84aa0d4 100644 --- a/src/geometry/euclidean/angle/oriented/right_angle.lean +++ b/src/geometry/euclidean/angle/oriented/right_angle.lean @@ -24,7 +24,7 @@ namespace orientation open finite_dimensional -variables {V : Type*} [inner_product_space ℝ V] +variables {V : Type*} [normed_add_comm_group V] [inner_product_space ℝ V] variables [hd2 : fact (finrank ℝ V = 2)] (o : orientation ℝ V (fin 2)) include hd2 o @@ -625,8 +625,9 @@ namespace euclidean_geometry open finite_dimensional -variables {V : Type*} {P : Type*} [inner_product_space ℝ V] [metric_space P] -variables [normed_add_torsor V P] [hd2 : fact (finrank ℝ V = 2)] [module.oriented ℝ V (fin 2)] +variables {V : Type*} {P : Type*} + [normed_add_comm_group V] [inner_product_space ℝ V] [metric_space P] [normed_add_torsor V P] + [hd2 : fact (finrank ℝ V = 2)] [module.oriented ℝ V (fin 2)] include hd2 /-- An angle in a right-angled triangle expressed using `arccos`. -/ diff --git a/src/geometry/euclidean/angle/oriented/rotation.lean b/src/geometry/euclidean/angle/oriented/rotation.lean index d0b86900c8488..0b2e8d1bbb0ee 100644 --- a/src/geometry/euclidean/angle/oriented/rotation.lean +++ b/src/geometry/euclidean/angle/oriented/rotation.lean @@ -26,7 +26,9 @@ namespace orientation local attribute [instance] fact_finite_dimensional_of_finrank_eq_succ local attribute [instance] complex.finrank_real_complex_fact -variables {V V' : Type*} [inner_product_space ℝ V] [inner_product_space ℝ V'] +variables {V V' : Type*} +variables [normed_add_comm_group V] [normed_add_comm_group V'] +variables [inner_product_space ℝ V] [inner_product_space ℝ V'] variables [fact (finrank ℝ V = 2)] [fact (finrank ℝ V' = 2)] (o : orientation ℝ V (fin 2)) local notation `J` := o.right_angle_rotation diff --git a/src/geometry/euclidean/angle/sphere.lean b/src/geometry/euclidean/angle/sphere.lean index 9293c24e9f52a..00f153eef699e 100644 --- a/src/geometry/euclidean/angle/sphere.lean +++ b/src/geometry/euclidean/angle/sphere.lean @@ -20,7 +20,7 @@ open_locale euclidean_geometry real real_inner_product_space complex_conjugate namespace orientation -variables {V : Type*} [inner_product_space ℝ V] +variables {V : Type*} [normed_add_comm_group V] [inner_product_space ℝ V] variables [fact (finrank ℝ V = 2)] (o : orientation ℝ V (fin 2)) /-- Angle at center of a circle equals twice angle at circumference, oriented vector angle @@ -67,8 +67,9 @@ end orientation namespace euclidean_geometry -variables {V : Type*} {P : Type*} [inner_product_space ℝ V] [metric_space P] -variables [normed_add_torsor V P] [hd2 : fact (finrank ℝ V = 2)] [module.oriented ℝ V (fin 2)] +variables {V : Type*} {P : Type*} + [normed_add_comm_group V] [inner_product_space ℝ V] [metric_space P] [normed_add_torsor V P] + [hd2 : fact (finrank ℝ V = 2)] [module.oriented ℝ V (fin 2)] include hd2 local notation `o` := module.oriented.positive_orientation @@ -248,8 +249,9 @@ namespace triangle open euclidean_geometry -variables {V : Type*} {P : Type*} [inner_product_space ℝ V] [metric_space P] -variables [normed_add_torsor V P] [hd2 : fact (finrank ℝ V = 2)] [module.oriented ℝ V (fin 2)] +variables {V : Type*} {P : Type*} + [normed_add_comm_group V] [inner_product_space ℝ V] [metric_space P] [normed_add_torsor V P] + [hd2 : fact (finrank ℝ V = 2)] [module.oriented ℝ V (fin 2)] include hd2 local notation `o` := module.oriented.positive_orientation @@ -353,8 +355,9 @@ end affine namespace euclidean_geometry -variables {V : Type*} {P : Type*} [inner_product_space ℝ V] [metric_space P] -variables [normed_add_torsor V P] [hd2 : fact (finrank ℝ V = 2)] [module.oriented ℝ V (fin 2)] +variables {V : Type*} {P : Type*} + [normed_add_comm_group V] [inner_product_space ℝ V] [metric_space P] [normed_add_torsor V P] + [hd2 : fact (finrank ℝ V = 2)] [module.oriented ℝ V (fin 2)] include hd2 local notation `o` := module.oriented.positive_orientation diff --git a/src/geometry/euclidean/angle/unoriented/affine.lean b/src/geometry/euclidean/angle/unoriented/affine.lean index 9a0025136b94f..8b4036ae625b1 100644 --- a/src/geometry/euclidean/angle/unoriented/affine.lean +++ b/src/geometry/euclidean/angle/unoriented/affine.lean @@ -27,8 +27,8 @@ namespace euclidean_geometry open inner_product_geometry -variables {V : Type*} {P : Type*} [inner_product_space ℝ V] [metric_space P] - [normed_add_torsor V P] +variables {V : Type*} {P : Type*} + [normed_add_comm_group V] [inner_product_space ℝ V] [metric_space P] [normed_add_torsor V P] include V /-- The undirected angle at `p2` between the line segments to `p1` and @@ -50,7 +50,8 @@ begin (continuous_snd.snd.vsub continuous_snd.fst)).continuous_at end -@[simp] lemma _root_.affine_isometry.angle_map {V₂ P₂ : Type*} [inner_product_space ℝ V₂] +@[simp] lemma _root_.affine_isometry.angle_map {V₂ P₂ : Type*} + [normed_add_comm_group V₂] [inner_product_space ℝ V₂] [metric_space P₂] [normed_add_torsor V₂ P₂] (f : P →ᵃⁱ[ℝ] P₂) (p₁ p₂ p₃ : P) : ∠ (f p₁) (f p₂) (f p₃) = ∠ p₁ p₂ p₃ := by simp_rw [angle, ←affine_isometry.map_vsub, linear_isometry.angle_map] diff --git a/src/geometry/euclidean/angle/unoriented/basic.lean b/src/geometry/euclidean/angle/unoriented/basic.lean index 0a3cc92f398bd..d332c358a4407 100644 --- a/src/geometry/euclidean/angle/unoriented/basic.lean +++ b/src/geometry/euclidean/angle/unoriented/basic.lean @@ -28,7 +28,7 @@ open_locale real_inner_product_space namespace inner_product_geometry -variables {V : Type*} [inner_product_space ℝ V] {x y : V} +variables {V : Type*} [normed_add_comm_group V] [inner_product_space ℝ V] {x y : V} /-- The undirected angle between two vectors. If either vector is 0, this is π/2. See `orientation.oangle` for the corresponding oriented angle @@ -48,6 +48,7 @@ by rw [angle, angle, real_inner_smul_left, inner_smul_right, norm_smul, norm_smu mul_mul_mul_comm _ (‖x‖), abs_mul_abs_self, ← mul_assoc c c, mul_div_mul_left _ _ this] @[simp] lemma _root_.linear_isometry.angle_map {E F : Type*} + [normed_add_comm_group E] [normed_add_comm_group F] [inner_product_space ℝ E] [inner_product_space ℝ F] (f : E →ₗᵢ[ℝ] F) (u v : E) : angle (f u) (f v) = angle u v := by rw [angle, angle, f.inner_map_map, f.norm_map, f.norm_map] @@ -112,7 +113,8 @@ end @[simp] lemma angle_self {x : V} (hx : x ≠ 0) : angle x x = 0 := begin unfold angle, - rw [←real_inner_self_eq_norm_mul_norm, div_self (inner_self_ne_zero.2 hx), real.arccos_one] + rw [←real_inner_self_eq_norm_mul_norm, div_self (inner_self_ne_zero.2 hx : ⟪x, x⟫ ≠ 0), + real.arccos_one] end /-- The angle between a nonzero vector and its negation. -/ diff --git a/src/geometry/euclidean/angle/unoriented/conformal.lean b/src/geometry/euclidean/angle/unoriented/conformal.lean index cd65fa3ca0463..d19eca5cf0dd6 100644 --- a/src/geometry/euclidean/angle/unoriented/conformal.lean +++ b/src/geometry/euclidean/angle/unoriented/conformal.lean @@ -15,11 +15,11 @@ This file proves that conformal maps preserve angles. namespace inner_product_geometry -variables {V : Type*} [inner_product_space ℝ V] +variables {E F : Type*} +variables [normed_add_comm_group E] [normed_add_comm_group F] +variables [inner_product_space ℝ E] [inner_product_space ℝ F] -lemma is_conformal_map.preserves_angle {E F : Type*} - [inner_product_space ℝ E] [inner_product_space ℝ F] - {f' : E →L[ℝ] F} (h : is_conformal_map f') (u v : E) : +lemma is_conformal_map.preserves_angle {f' : E →L[ℝ] F} (h : is_conformal_map f') (u v : E) : angle (f' u) (f' v) = angle u v := begin obtain ⟨c, hc, li, rfl⟩ := h, @@ -28,9 +28,7 @@ end /-- If a real differentiable map `f` is conformal at a point `x`, then it preserves the angles at that point. -/ -lemma conformal_at.preserves_angle {E F : Type*} - [inner_product_space ℝ E] [inner_product_space ℝ F] - {f : E → F} {x : E} {f' : E →L[ℝ] F} +lemma conformal_at.preserves_angle {f : E → F} {x : E} {f' : E →L[ℝ] F} (h : has_fderiv_at f f' x) (H : conformal_at f x) (u v : E) : angle (f' u) (f' v) = angle u v := let ⟨f₁, h₁, c⟩ := H in h₁.unique h ▸ is_conformal_map.preserves_angle c u v diff --git a/src/geometry/euclidean/angle/unoriented/right_angle.lean b/src/geometry/euclidean/angle/unoriented/right_angle.lean index 8c6a7b71c02e8..75beb298c6074 100644 --- a/src/geometry/euclidean/angle/unoriented/right_angle.lean +++ b/src/geometry/euclidean/angle/unoriented/right_angle.lean @@ -32,7 +32,7 @@ open_locale real_inner_product_space namespace inner_product_geometry -variables {V : Type*} [inner_product_space ℝ V] +variables {V : Type*} [normed_add_comm_group V] [inner_product_space ℝ V] /-- Pythagorean theorem, if-and-only-if vector angle form. -/ lemma norm_add_sq_eq_norm_sq_add_norm_sq_iff_angle_eq_pi_div_two (x y : V) : @@ -387,8 +387,8 @@ namespace euclidean_geometry open inner_product_geometry -variables {V : Type*} {P : Type*} [inner_product_space ℝ V] [metric_space P] - [normed_add_torsor V P] +variables {V : Type*} {P : Type*} + [normed_add_comm_group V] [inner_product_space ℝ V] [metric_space P] [normed_add_torsor V P] include V /-- **Pythagorean theorem**, if-and-only-if angle-at-point form. -/ diff --git a/src/geometry/euclidean/basic.lean b/src/geometry/euclidean/basic.lean index f2e3bb48c6317..86f03e5a75ac3 100644 --- a/src/geometry/euclidean/basic.lean +++ b/src/geometry/euclidean/basic.lean @@ -27,8 +27,9 @@ proofs or more geometrical content generally go in separate files. ## Implementation notes To declare `P` as the type of points in a Euclidean affine space with -`V` as the type of vectors, use `[inner_product_space ℝ V] [metric_space P] -[normed_add_torsor V P]`. This works better with `out_param` to make +`V` as the type of vectors, use +`[normed_add_comm_group V] [inner_product_space ℝ V] [metric_space P] [normed_add_torsor V P]`. +This works better with `out_param` to make `V` implicit in most cases than having a separate type alias for Euclidean affine spaces. @@ -55,8 +56,9 @@ This section develops some geometrical definitions and results on Euclidean affine spaces. -/ -variables {V : Type*} {P : Type*} [inner_product_space ℝ V] [metric_space P] - [normed_add_torsor V P] +variables {V : Type*} {P : Type*} +variables [normed_add_comm_group V] [inner_product_space ℝ V] [metric_space P] +variables [normed_add_torsor V P] include V /-- The midpoint of the segment AB is the same distance from A as it is from B. -/ @@ -89,7 +91,7 @@ lemma dist_affine_combination {ι : Type*} {s : finset ι} {w₁ w₂ : ι → (w₁ - w₂) i₁ * (w₁ - w₂) i₂ * (dist (p i₁) (p i₂) * dist (p i₁) (p i₂))) / 2 := begin rw [dist_eq_norm_vsub V (s.affine_combination p w₁) (s.affine_combination p w₂), - ←inner_self_eq_norm_mul_norm, finset.affine_combination_vsub], + ←@inner_self_eq_norm_mul_norm ℝ, finset.affine_combination_vsub], have h : ∑ i in s, (w₁ - w₂) i = 0, { simp_rw [pi.sub_apply, finset.sum_sub_distrib, h₁, h₂, sub_self] }, exact inner_weighted_vsub p h p h diff --git a/src/geometry/euclidean/circumcenter.lean b/src/geometry/euclidean/circumcenter.lean index b5250e2d515f3..e25e36dc17fa7 100644 --- a/src/geometry/euclidean/circumcenter.lean +++ b/src/geometry/euclidean/circumcenter.lean @@ -34,8 +34,8 @@ open_locale real_inner_product_space namespace euclidean_geometry -variables {V : Type*} {P : Type*} [inner_product_space ℝ V] [metric_space P] - [normed_add_torsor V P] +variables {V : Type*} {P : Type*} + [normed_add_comm_group V] [inner_product_space ℝ V] [metric_space P] [normed_add_torsor V P] include V open affine_subspace @@ -252,8 +252,8 @@ namespace simplex open finset affine_subspace euclidean_geometry -variables {V : Type*} {P : Type*} [inner_product_space ℝ V] [metric_space P] - [normed_add_torsor V P] +variables {V : Type*} {P : Type*} + [normed_add_comm_group V] [inner_product_space ℝ V] [metric_space P] [normed_add_torsor V P] include V /-- The circumsphere of a simplex. -/ @@ -735,8 +735,8 @@ namespace euclidean_geometry open affine affine_subspace finite_dimensional -variables {V : Type*} {P : Type*} [inner_product_space ℝ V] [metric_space P] - [normed_add_torsor V P] +variables {V : Type*} {P : Type*} + [normed_add_comm_group V] [inner_product_space ℝ V] [metric_space P] [normed_add_torsor V P] include V /-- Given a nonempty affine subspace, whose direction is complete, diff --git a/src/geometry/euclidean/inversion.lean b/src/geometry/euclidean/inversion.lean index e16f28592db29..ff0d2dcf46fec 100644 --- a/src/geometry/euclidean/inversion.lean +++ b/src/geometry/euclidean/inversion.lean @@ -25,7 +25,8 @@ open metric real function namespace euclidean_geometry -variables {V P : Type*} [inner_product_space ℝ V] [metric_space P] [normed_add_torsor V P] +variables {V P : Type*} + [normed_add_comm_group V] [inner_product_space ℝ V] [metric_space P] [normed_add_torsor V P] {a b c d x y z : P} {R : ℝ} include V diff --git a/src/geometry/euclidean/monge_point.lean b/src/geometry/euclidean/monge_point.lean index b5713ba092a15..981a00e57a85e 100644 --- a/src/geometry/euclidean/monge_point.lean +++ b/src/geometry/euclidean/monge_point.lean @@ -57,8 +57,8 @@ namespace simplex open finset affine_subspace euclidean_geometry points_with_circumcenter_index -variables {V : Type*} {P : Type*} [inner_product_space ℝ V] [metric_space P] - [normed_add_torsor V P] +variables {V : Type*} {P : Type*} + [normed_add_comm_group V] [inner_product_space ℝ V] [metric_space P] [normed_add_torsor V P] include V /-- The Monge point of a simplex (in 2 or more dimensions) is a @@ -443,8 +443,8 @@ namespace triangle open euclidean_geometry finset simplex affine_subspace finite_dimensional -variables {V : Type*} {P : Type*} [inner_product_space ℝ V] [metric_space P] - [normed_add_torsor V P] +variables {V : Type*} {P : Type*} + [normed_add_comm_group V] [inner_product_space ℝ V] [metric_space P] [normed_add_torsor V P] include V /-- The orthocenter of a triangle is the intersection of its @@ -626,8 +626,8 @@ namespace euclidean_geometry open affine affine_subspace finite_dimensional -variables {V : Type*} {P : Type*} [inner_product_space ℝ V] [metric_space P] - [normed_add_torsor V P] +variables {V : Type*} {P : Type*} + [normed_add_comm_group V] [inner_product_space ℝ V] [metric_space P] [normed_add_torsor V P] include V diff --git a/src/geometry/euclidean/sphere/basic.lean b/src/geometry/euclidean/sphere/basic.lean index 91372a3c373a7..a017fe8d65a47 100644 --- a/src/geometry/euclidean/sphere/basic.lean +++ b/src/geometry/euclidean/sphere/basic.lean @@ -185,7 +185,8 @@ lemma concyclic_pair (p₁ p₂ : P) : concyclic ({p₁, p₂} : set P) := end normed_space section euclidean_space -variables [inner_product_space ℝ V] [metric_space P] [normed_add_torsor V P] +variables + [normed_add_comm_group V] [inner_product_space ℝ V] [metric_space P] [normed_add_torsor V P] include V /-- Any three points in a cospherical set are affinely independent. -/ diff --git a/src/geometry/euclidean/sphere/power.lean b/src/geometry/euclidean/sphere/power.lean index 4a01046f9d995..ad14d1991684f 100644 --- a/src/geometry/euclidean/sphere/power.lean +++ b/src/geometry/euclidean/sphere/power.lean @@ -21,7 +21,7 @@ secants) in spheres in real inner product spaces and Euclidean affine spaces. open real open_locale euclidean_geometry real_inner_product_space real -variables {V : Type*} [inner_product_space ℝ V] +variables {V : Type*} [normed_add_comm_group V] [inner_product_space ℝ V] namespace inner_product_geometry diff --git a/src/geometry/euclidean/sphere/ptolemy.lean b/src/geometry/euclidean/sphere/ptolemy.lean index cf45fc41b4cdc..aa5d6eaa4fec3 100644 --- a/src/geometry/euclidean/sphere/ptolemy.lean +++ b/src/geometry/euclidean/sphere/ptolemy.lean @@ -42,7 +42,7 @@ open_locale euclidean_geometry real_inner_product_space real namespace euclidean_geometry -variables {V : Type*} [inner_product_space ℝ V] +variables {V : Type*} [normed_add_comm_group V] [inner_product_space ℝ V] variables {P : Type*} [metric_space P] [normed_add_torsor V P] include V diff --git a/src/geometry/euclidean/sphere/second_inter.lean b/src/geometry/euclidean/sphere/second_inter.lean index e00f064381610..ec52336826a8c 100644 --- a/src/geometry/euclidean/sphere/second_inter.lean +++ b/src/geometry/euclidean/sphere/second_inter.lean @@ -23,7 +23,8 @@ open_locale real_inner_product_space namespace euclidean_geometry -variables {V : Type*} {P : Type*} [inner_product_space ℝ V] [metric_space P] [normed_add_torsor V P] +variables {V : Type*} {P : Type*} + [normed_add_comm_group V] [inner_product_space ℝ V] [metric_space P] [normed_add_torsor V P] include V /-- The second intersection of a sphere with a line through a point on that sphere; that point diff --git a/src/geometry/euclidean/triangle.lean b/src/geometry/euclidean/triangle.lean index c97326e4b4241..7e8e292c870a9 100644 --- a/src/geometry/euclidean/triangle.lean +++ b/src/geometry/euclidean/triangle.lean @@ -50,7 +50,7 @@ most conveniently be developed in terms of vectors and then used to deduce corresponding results for Euclidean affine spaces. -/ -variables {V : Type*} [inner_product_space ℝ V] +variables {V : Type*} [normed_add_comm_group V] [inner_product_space ℝ V] /-- Law of cosines (cosine rule), vector angle form. -/ lemma norm_sub_sq_eq_norm_sq_add_norm_sq_sub_two_mul_norm_mul_norm_mul_cos_angle @@ -262,8 +262,8 @@ open inner_product_geometry open_locale euclidean_geometry -variables {V : Type*} {P : Type*} [inner_product_space ℝ V] [metric_space P] - [normed_add_torsor V P] +variables {V : Type*} {P : Type*} + [normed_add_comm_group V] [inner_product_space ℝ V] [metric_space P] [normed_add_torsor V P] include V /-- **Law of cosines** (cosine rule), angle-at-point form. -/ diff --git a/src/geometry/manifold/instances/sphere.lean b/src/geometry/manifold/instances/sphere.lean index 113a65c86380f..c8ad586833bae 100644 --- a/src/geometry/manifold/instances/sphere.lean +++ b/src/geometry/manifold/instances/sphere.lean @@ -55,7 +55,7 @@ naive expression `euclidean_space ℝ (fin (finrank ℝ E - 1))` for the model s `euclidean_space ℝ (fin 1)`. -/ -variables {E : Type*} [inner_product_space ℝ E] +variables {E : Type*} [normed_add_comm_group E] [inner_product_space ℝ E] noncomputable theory @@ -95,7 +95,7 @@ end lemma continuous_on_stereo_to_fun [complete_space E] : continuous_on (stereo_to_fun v) {x : E | innerSL _ v x ≠ (1:ℝ)} := -(@cont_diff_on_stereo_to_fun E _ v _).continuous_on +(@cont_diff_on_stereo_to_fun E _ _ v _).continuous_on variables (v) @@ -162,7 +162,7 @@ end lemma cont_diff_stereo_inv_fun_aux : cont_diff ℝ ⊤ (stereo_inv_fun_aux v) := begin - have h₀ : cont_diff ℝ ⊤ (λ w : E, ‖w‖ ^ 2) := cont_diff_norm_sq, + have h₀ : cont_diff ℝ ⊤ (λ w : E, ‖w‖ ^ 2) := cont_diff_norm_sq ℝ, have h₁ : cont_diff ℝ ⊤ (λ w : E, (‖w‖ ^ 2 + 4)⁻¹), { refine (h₀.add cont_diff_const).inv _, intros x, @@ -417,7 +417,7 @@ begin split, { exact continuous_subtype_coe }, { intros v _, - let U := -- Again, removing type ascription... + let U : _ ≃ₗᵢ[ℝ] _ := -- Again, partially removing type ascription... (orthonormal_basis.from_orthogonal_span_singleton n (ne_zero_of_mem_unit_sphere (-v))).repr, exact ((cont_diff_stereo_inv_fun_aux.comp @@ -438,7 +438,7 @@ begin rw cont_mdiff_iff_target, refine ⟨continuous_induced_rng.2 hf.continuous, _⟩, intros v, - let U := -- Again, removing type ascription... Weird that this helps! + let U : _ ≃ₗᵢ[ℝ] _ := -- Again, partially removing type ascription... Weird that this helps! (orthonormal_basis.from_orthogonal_span_singleton n (ne_zero_of_mem_unit_sphere (-v))).repr, have h : cont_diff_on ℝ ⊤ _ set.univ := U.cont_diff.cont_diff_on, diff --git a/src/linear_algebra/matrix/ldl.lean b/src/linear_algebra/matrix/ldl.lean index d21b4263db629..f346a739ba6b3 100644 --- a/src/linear_algebra/matrix/ldl.lean +++ b/src/linear_algebra/matrix/ldl.lean @@ -41,11 +41,13 @@ applying Gram-Schmidt-Orthogonalization w.r.t. the inner product induced by `S basis vectors `pi.basis_fun`. -/ noncomputable def LDL.lower_inv : matrix n n 𝕜 := @gram_schmidt - 𝕜 (n → 𝕜) _ (inner_product_space.of_matrix hS.transpose) n _ _ _ (pi.basis_fun 𝕜 n) + 𝕜 (n → 𝕜) _ + (_ : _) + (inner_product_space.of_matrix hS.transpose) n _ _ _ (pi.basis_fun 𝕜 n) lemma LDL.lower_inv_eq_gram_schmidt_basis : LDL.lower_inv hS = ((pi.basis_fun 𝕜 n).to_matrix - (@gram_schmidt_basis 𝕜 (n → 𝕜) _ + (@gram_schmidt_basis 𝕜 (n → 𝕜) _ (_ : _) (inner_product_space.of_matrix hS.transpose) n _ _ _ (pi.basis_fun 𝕜 n)))ᵀ := begin ext i j, @@ -57,17 +59,14 @@ noncomputable instance LDL.invertible_lower_inv : invertible (LDL.lower_inv hS) begin rw [LDL.lower_inv_eq_gram_schmidt_basis], haveI := basis.invertible_to_matrix (pi.basis_fun 𝕜 n) - (@gram_schmidt_basis 𝕜 (n → 𝕜) _ (inner_product_space.of_matrix hS.transpose) + (@gram_schmidt_basis 𝕜 (n → 𝕜) _ (_ : _) (inner_product_space.of_matrix hS.transpose) n _ _ _ (pi.basis_fun 𝕜 n)), apply_instance end lemma LDL.lower_inv_orthogonal {i j : n} (h₀ : i ≠ j) : ⟪(LDL.lower_inv hS i), Sᵀ.mul_vec (LDL.lower_inv hS j)⟫ₑ = 0 := -show @inner 𝕜 (n → 𝕜) (inner_product_space.of_matrix hS.transpose).to_has_inner - (LDL.lower_inv hS i) - (LDL.lower_inv hS j) = 0, -by apply gram_schmidt_orthogonal _ _ h₀ +@gram_schmidt_orthogonal 𝕜 _ _ (_ : _) (inner_product_space.of_matrix hS.transpose) _ _ _ _ _ _ _ h₀ /-- The entries of the diagonal matrix `D` of the LDL decomposition. -/ noncomputable def LDL.diag_entries : n → 𝕜 := @@ -78,7 +77,8 @@ noncomputable def LDL.diag : matrix n n 𝕜 := matrix.diagonal (LDL.diag_entrie lemma LDL.lower_inv_triangular {i j : n} (hij : i < j) : LDL.lower_inv hS i j = 0 := -by rw [← @gram_schmidt_triangular 𝕜 (n → 𝕜) _ (inner_product_space.of_matrix hS.transpose) n _ _ _ +by rw [← @gram_schmidt_triangular + 𝕜 (n → 𝕜) _ (_ : _) (inner_product_space.of_matrix hS.transpose) n _ _ _ i j hij (pi.basis_fun 𝕜 n), pi.basis_fun_repr, LDL.lower_inv] /-- Inverse statement of **LDL decomposition**: we can conjugate a positive definite matrix diff --git a/src/linear_algebra/matrix/pos_def.lean b/src/linear_algebra/matrix/pos_def.lean index cb182b91d8f2b..c814a4996c235 100644 --- a/src/linear_algebra/matrix/pos_def.lean +++ b/src/linear_algebra/matrix/pos_def.lean @@ -139,10 +139,11 @@ namespace matrix variables {𝕜 : Type*} [is_R_or_C 𝕜] {n : Type*} [fintype n] -/-- A positive definite matrix `M` induces an inner product `⟪x, y⟫ = xᴴMy`. -/ -noncomputable def inner_product_space.of_matrix - {M : matrix n n 𝕜} (hM : M.pos_def) : inner_product_space 𝕜 (n → 𝕜) := -inner_product_space.of_core +/-- A positive definite matrix `M` induces a norm `‖x‖ = sqrt (re xᴴMx)`. -/ +@[reducible] +noncomputable def normed_add_comm_group.of_matrix {M : matrix n n 𝕜} (hM : M.pos_def) : + normed_add_comm_group (n → 𝕜) := +@inner_product_space.of_core.to_normed_add_comm_group _ _ _ _ _ { inner := λ x y, dot_product (star x) (M.mul_vec y), conj_symm := λ x y, by rw [star_dot_product, star_ring_end_apply, star_star, star_mul_vec, @@ -161,4 +162,9 @@ inner_product_space.of_core add_left := by simp only [star_add, add_dot_product, eq_self_iff_true, forall_const], smul_left := λ x y r, by rw [← smul_eq_mul, ←smul_dot_product, star_ring_end_apply, ← star_smul] } +/-- A positive definite matrix `M` induces an inner product `⟪x, y⟫ = xᴴMy`. -/ +def inner_product_space.of_matrix {M : matrix n n 𝕜} (hM : M.pos_def) : + @inner_product_space 𝕜 (n → 𝕜) _ (normed_add_comm_group.of_matrix hM) := +inner_product_space.of_core _ + end matrix diff --git a/src/linear_algebra/matrix/spectrum.lean b/src/linear_algebra/matrix/spectrum.lean index 632d74e1e7e9b..0d533048e00fe 100644 --- a/src/linear_algebra/matrix/spectrum.lean +++ b/src/linear_algebra/matrix/spectrum.lean @@ -99,8 +99,8 @@ begin simp_rw [mul_vec_single, mul_one, orthonormal_basis.coe_to_basis_repr_apply, orthonormal_basis.repr_reindex], refl }, - { simp only [diagonal_mul, (∘), eigenvalues, eigenvector_basis], - rw [basis.to_matrix_apply, + { simp only [diagonal_mul, (∘), eigenvalues], + rw [eigenvector_basis, basis.to_matrix_apply, orthonormal_basis.coe_to_basis_repr_apply, orthonormal_basis.repr_reindex, eigenvalues₀, pi_Lp.basis_fun_apply, pi_Lp.equiv_symm_single] } end diff --git a/src/measure_theory/function/ae_eq_of_integral.lean b/src/measure_theory/function/ae_eq_of_integral.lean index c83b0b93f2156..9ab0cfcf0be01 100644 --- a/src/measure_theory/function/ae_eq_of_integral.lean +++ b/src/measure_theory/function/ae_eq_of_integral.lean @@ -52,7 +52,8 @@ section ae_eq_of_forall variables {α E 𝕜 : Type*} {m : measurable_space α} {μ : measure α} [is_R_or_C 𝕜] -lemma ae_eq_zero_of_forall_inner [inner_product_space 𝕜 E] [second_countable_topology E] +lemma ae_eq_zero_of_forall_inner + [normed_add_comm_group E] [inner_product_space 𝕜 E] [second_countable_topology E] {f : α → E} (hf : ∀ c : E, (λ x, (inner c (f x) : 𝕜)) =ᵐ[μ] 0) : f =ᵐ[μ] 0 := begin @@ -60,7 +61,7 @@ begin have hs : dense_range s := dense_range_dense_seq E, have hf' : ∀ᵐ x ∂μ, ∀ n : ℕ, inner (s n) (f x) = (0 : 𝕜), from ae_all_iff.mpr (λ n, hf (s n)), refine hf'.mono (λ x hx, _), - rw [pi.zero_apply, ← inner_self_eq_zero], + rw [pi.zero_apply, ← @inner_self_eq_zero 𝕜], have h_closed : is_closed {c : E | inner c (f x) = (0 : 𝕜)}, from is_closed_eq (continuous_id.inner continuous_const) continuous_const, exact @is_closed_property ℕ E _ s (λ c, inner c (f x) = (0 : 𝕜)) hs h_closed (λ n, hx n) _, diff --git a/src/measure_theory/function/conditional_expectation/basic.lean b/src/measure_theory/function/conditional_expectation/basic.lean index 46a556bb163ba..0cacb74285b49 100644 --- a/src/measure_theory/function/conditional_expectation/basic.lean +++ b/src/measure_theory/function/conditional_expectation/basic.lean @@ -137,7 +137,7 @@ begin exact eventually_eq.fun_comp hff' (λ x, c • x), end -lemma const_inner {𝕜 β} [is_R_or_C 𝕜] [inner_product_space 𝕜 β] +lemma const_inner {𝕜 β} [is_R_or_C 𝕜] [normed_add_comm_group β] [inner_product_space 𝕜 β] {f : α → β} (hfm : ae_strongly_measurable' m f μ) (c : β) : ae_strongly_measurable' m (λ x, (inner c (f x) : 𝕜)) μ := begin @@ -218,9 +218,9 @@ variables {α β γ E E' F F' G G' H 𝕜 : Type*} {p : ℝ≥0∞} [is_R_or_C 𝕜] -- 𝕜 for ℝ or ℂ [topological_space β] -- β for a generic topological space -- E for an inner product space - [inner_product_space 𝕜 E] + [normed_add_comm_group E] [inner_product_space 𝕜 E] -- E' for an inner product space on which we compute integrals - [inner_product_space 𝕜 E'] + [normed_add_comm_group E'] [inner_product_space 𝕜 E'] [complete_space E'] [normed_space ℝ E'] -- F for a Lp submodule [normed_add_comm_group F] [normed_space 𝕜 F] @@ -754,6 +754,7 @@ begin end include 𝕜 +variables (𝕜) lemma Lp.ae_eq_zero_of_forall_set_integral_eq_zero' (hm : m ≤ m0) (f : Lp E' p μ) (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) @@ -798,10 +799,11 @@ begin exact (hf_int_finite s hs hμs).sub (hg_int_finite s hs hμs), }, have hfg_meas : ae_strongly_measurable' m ⇑(f - g) μ, from ae_strongly_measurable'.congr (hf_meas.sub hg_meas) (Lp.coe_fn_sub f g).symm, - exact Lp.ae_eq_zero_of_forall_set_integral_eq_zero' hm (f-g) hp_ne_zero hp_ne_top hfg_int hfg' + exact Lp.ae_eq_zero_of_forall_set_integral_eq_zero' 𝕜 hm (f-g) hp_ne_zero hp_ne_top hfg_int hfg' hfg_meas, end +variables {𝕜} omit 𝕜 lemma ae_eq_of_forall_set_integral_eq_of_sigma_finite' (hm : m ≤ m0) [sigma_finite (μ.trim hm)] @@ -917,7 +919,7 @@ local notation `⟪`x`, `y`⟫₂` := @inner 𝕜 (α →₂[μ] E) _ x y variables (𝕜) /-- Conditional expectation of a function in L2 with respect to a sigma-algebra -/ def condexp_L2 (hm : m ≤ m0) : (α →₂[μ] E) →L[𝕜] (Lp_meas E 𝕜 m 2 μ) := -@orthogonal_projection 𝕜 (α →₂[μ] E) _ _ (Lp_meas E 𝕜 m 2 μ) +@orthogonal_projection 𝕜 (α →₂[μ] E) _ _ _ (Lp_meas E 𝕜 m 2 μ) (by { haveI : fact (m ≤ m0) := ⟨hm⟩, exact infer_instance, }) variables {𝕜} @@ -935,11 +937,11 @@ lemma integrable_condexp_L2_of_is_finite_measure (hm : m ≤ m0) [is_finite_meas integrable (condexp_L2 𝕜 hm f) μ := integrable_on_univ.mp $ integrable_on_condexp_L2_of_measure_ne_top hm (measure_ne_top _ _) f -lemma norm_condexp_L2_le_one (hm : m ≤ m0) : ‖@condexp_L2 α E 𝕜 _ _ _ _ _ μ hm‖ ≤ 1 := +lemma norm_condexp_L2_le_one (hm : m ≤ m0) : ‖@condexp_L2 α E 𝕜 _ _ _ _ _ _ μ hm‖ ≤ 1 := by { haveI : fact (m ≤ m0) := ⟨hm⟩, exact orthogonal_projection_norm_le _, } lemma norm_condexp_L2_le (hm : m ≤ m0) (f : α →₂[μ] E) : ‖condexp_L2 𝕜 hm f‖ ≤ ‖f‖ := -((@condexp_L2 _ E 𝕜 _ _ _ _ _ μ hm).le_op_norm f).trans +((@condexp_L2 _ E 𝕜 _ _ _ _ _ _ μ hm).le_op_norm f).trans (mul_le_of_le_one_left (norm_nonneg _) (norm_condexp_L2_le_one hm)) lemma snorm_condexp_L2_le (hm : m ≤ m0) (f : α →₂[μ] E) : @@ -1081,7 +1083,7 @@ begin { refine mem_ℒp.const_inner _ _, rw Lp_meas_coe, exact Lp.mem_ℒp _, }, have h_eq : h_mem_Lp.to_Lp _ =ᵐ[μ] λ a, ⟪c, condexp_L2 𝕜 hm f a⟫, from h_mem_Lp.coe_fn_to_Lp, refine eventually_eq.trans _ h_eq, - refine Lp.ae_eq_of_forall_set_integral_eq' hm _ _ two_ne_zero ennreal.coe_ne_top + refine Lp.ae_eq_of_forall_set_integral_eq' 𝕜 hm _ _ two_ne_zero ennreal.coe_ne_top (λ s hs hμs, integrable_on_condexp_L2_of_measure_ne_top hm hμs.ne _) _ _ _ _, { intros s hs hμs, rw [integrable_on, integrable_congr (ae_restrict_of_ae h_eq)], @@ -1107,7 +1109,7 @@ begin rw [← sub_eq_zero, Lp_meas_coe, ← integral_sub' (integrable_on_Lp_of_measure_ne_top _ fact_one_le_two_ennreal.elim hμs) (integrable_on_Lp_of_measure_ne_top _ fact_one_le_two_ennreal.elim hμs)], - refine integral_eq_zero_of_forall_integral_inner_eq_zero _ _ _, + refine integral_eq_zero_of_forall_integral_inner_eq_zero 𝕜 _ _ _, { rw integrable_congr (ae_restrict_of_ae (Lp.coe_fn_sub ↑(condexp_L2 𝕜 hm f) f).symm), exact integrable_on_Lp_of_measure_ne_top _ fact_one_le_two_ennreal.elim hμs, }, intro c, @@ -1122,14 +1124,14 @@ begin exact integral_condexp_L2_eq_of_fin_meas_real _ hs hμs, end -variables {E'' 𝕜' : Type*} [is_R_or_C 𝕜'] +variables {E'' 𝕜' : Type*} [is_R_or_C 𝕜'] [normed_add_comm_group E''] [inner_product_space 𝕜' E''] [complete_space E''] [normed_space ℝ E''] variables (𝕜 𝕜') lemma condexp_L2_comp_continuous_linear_map (hm : m ≤ m0) (T : E' →L[ℝ] E'') (f : α →₂[μ] E') : (condexp_L2 𝕜' hm (T.comp_Lp f) : α →₂[μ] E'') =ᵐ[μ] T.comp_Lp (condexp_L2 𝕜 hm f : α →₂[μ] E') := begin - refine Lp.ae_eq_of_forall_set_integral_eq' hm _ _ two_ne_zero ennreal.coe_ne_top + refine Lp.ae_eq_of_forall_set_integral_eq' 𝕜' hm _ _ two_ne_zero ennreal.coe_ne_top (λ s hs hμs, integrable_on_condexp_L2_of_measure_ne_top hm hμs.ne _) (λ s hs hμs, integrable_on_Lp_of_measure_ne_top _ fact_one_le_two_ennreal.elim hμs.ne) _ _ _, @@ -1321,7 +1323,8 @@ lemma set_integral_condexp_L2_indicator (hs : measurable_set[m] s) (ht : measura ∫ x in s, (condexp_L2 ℝ hm (indicator_const_Lp 2 ht hμt (1 : ℝ))) x ∂μ = (μ (t ∩ s)).to_real := calc ∫ x in s, (condexp_L2 ℝ hm (indicator_const_Lp 2 ht hμt (1 : ℝ))) x ∂μ = ∫ x in s, indicator_const_Lp 2 ht hμt (1 : ℝ) x ∂μ : - @integral_condexp_L2_eq α _ ℝ _ _ _ _ _ _ _ _ hm (indicator_const_Lp 2 ht hμt (1 : ℝ)) hs hμs + @integral_condexp_L2_eq + α _ ℝ _ _ _ _ _ _ _ _ _ hm (indicator_const_Lp 2 ht hμt (1 : ℝ)) hs hμs ... = (μ (t ∩ s)).to_real • 1 : set_integral_indicator_const_Lp (hm s hs) ht hμt (1 : ℝ) ... = (μ (t ∩ s)).to_real : by rw [smul_eq_mul, mul_one] diff --git a/src/measure_theory/function/l1_space.lean b/src/measure_theory/function/l1_space.lean index d7b31184c0ecc..2f1c344e116b0 100644 --- a/src/measure_theory/function/l1_space.lean +++ b/src/measure_theory/function/l1_space.lean @@ -1014,7 +1014,8 @@ by { rw ← mem_ℒp_one_iff_integrable at hf ⊢, exact hf.im, } end is_R_or_C section inner_product -variables {𝕜 E : Type*} [is_R_or_C 𝕜] [inner_product_space 𝕜 E] {f : α → E} +variables {𝕜 E : Type*} +variables [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {f : α → E} local notation `⟪`x`, `y`⟫` := @inner 𝕜 E _ x y diff --git a/src/measure_theory/function/l2_space.lean b/src/measure_theory/function/l2_space.lean index a4158508000d6..b7b913670cbc8 100644 --- a/src/measure_theory/function/l2_space.lean +++ b/src/measure_theory/function/l2_space.lean @@ -58,7 +58,7 @@ end namespace L2 variables {α E F 𝕜 : Type*} [is_R_or_C 𝕜] [measurable_space α] {μ : measure α} - [inner_product_space 𝕜 E] [normed_add_comm_group F] + [normed_add_comm_group E] [inner_product_space 𝕜 E] [normed_add_comm_group F] local notation `⟪`x`, `y`⟫` := @inner 𝕜 _ _ x y diff --git a/src/measure_theory/function/lp_space.lean b/src/measure_theory/function/lp_space.lean index f26bde6cd752d..b4b9c72fdc412 100644 --- a/src/measure_theory/function/lp_space.lean +++ b/src/measure_theory/function/lp_space.lean @@ -1418,7 +1418,7 @@ end end is_R_or_C section inner_product -variables {E' 𝕜 : Type*} [is_R_or_C 𝕜] [inner_product_space 𝕜 E'] +variables {E' 𝕜 : Type*} [is_R_or_C 𝕜] [normed_add_comm_group E'] [inner_product_space 𝕜 E'] local notation `⟪`x`, `y`⟫` := @inner 𝕜 E' _ x y diff --git a/src/measure_theory/function/special_functions/inner.lean b/src/measure_theory/function/special_functions/inner.lean index 097b21c6f8db4..627872248de06 100644 --- a/src/measure_theory/function/special_functions/inner.lean +++ b/src/measure_theory/function/special_functions/inner.lean @@ -11,7 +11,8 @@ import measure_theory.constructions.borel_space # Measurability of scalar products -/ -variables {α : Type*} {𝕜 : Type*} {E : Type*} [is_R_or_C 𝕜] [inner_product_space 𝕜 E] +variables {α : Type*} {𝕜 : Type*} {E : Type*} +variables [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] local notation `⟪`x`, `y`⟫` := @inner 𝕜 _ _ x y @[measurability] diff --git a/src/measure_theory/function/strongly_measurable/inner.lean b/src/measure_theory/function/strongly_measurable/inner.lean index 0795426893125..6b25e4682327c 100644 --- a/src/measure_theory/function/strongly_measurable/inner.lean +++ b/src/measure_theory/function/strongly_measurable/inner.lean @@ -18,7 +18,8 @@ namespace measure_theory namespace strongly_measurable -protected lemma inner {𝕜 : Type*} {E : Type*} [is_R_or_C 𝕜] [inner_product_space 𝕜 E] +protected lemma inner {𝕜 : Type*} {E : Type*} + [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {m : measurable_space α} {f g : α → E} (hf : strongly_measurable f) (hg : strongly_measurable g) : strongly_measurable (λ t, @inner 𝕜 _ _(f t) (g t)) := continuous.comp_strongly_measurable continuous_inner (hf.prod_mk hg) @@ -28,7 +29,7 @@ end strongly_measurable namespace ae_strongly_measurable variables {m : measurable_space α} {μ : measure α} {𝕜 : Type*} {E : Type*} [is_R_or_C 𝕜] - [inner_product_space 𝕜 E] + [normed_add_comm_group E] [inner_product_space 𝕜 E] local notation `⟪`x`, `y`⟫` := @inner 𝕜 _ _ x y protected lemma re {f : α → 𝕜} (hf : ae_strongly_measurable f μ) : diff --git a/src/measure_theory/integral/set_integral.lean b/src/measure_theory/integral/set_integral.lean index 3e0114cb2705e..85908862cbf86 100644 --- a/src/measure_theory/integral/set_integral.lean +++ b/src/measure_theory/integral/set_integral.lean @@ -1126,7 +1126,9 @@ end section inner -variables {E' : Type*} [inner_product_space 𝕜 E'] [complete_space E'] [normed_space ℝ E'] +variables {E' : Type*} +variables [normed_add_comm_group E'] [inner_product_space 𝕜 E'] +variables [complete_space E'] [normed_space ℝ E'] local notation `⟪`x`, `y`⟫` := @inner 𝕜 E' _ x y @@ -1134,6 +1136,10 @@ lemma integral_inner {f : α → E'} (hf : integrable f μ) (c : E') : ∫ x, ⟪c, f x⟫ ∂μ = ⟪c, ∫ x, f x ∂μ⟫ := ((innerSL 𝕜 c).restrict_scalars ℝ).integral_comp_comm hf +variables (𝕜) +-- variable binder update doesn't work for lemmas which refer to `𝕜` only via the notation +local notation (name := inner_with_explicit) `⟪`x`, `y`⟫` := @inner 𝕜 E' _ x y + lemma integral_eq_zero_of_forall_integral_inner_eq_zero (f : α → E') (hf : integrable f μ) (hf_int : ∀ (c : E'), ∫ x, ⟪c, f x⟫ ∂μ = 0) : ∫ x, f x ∂μ = 0 := diff --git a/src/measure_theory/measure/haar_of_basis.lean b/src/measure_theory/measure/haar_of_basis.lean index e8da34d00525d..ae4878316a13c 100644 --- a/src/measure_theory/measure/haar_of_basis.lean +++ b/src/measure_theory/measure/haar_of_basis.lean @@ -151,7 +151,8 @@ volume `1` to the parallelepiped spanned by any orthonormal basis. We define the some arbitrary choice of orthonormal basis. The fact that it works with any orthonormal basis is proved in `orthonormal_basis.volume_parallelepiped`. -/ @[priority 100] instance measure_space_of_inner_product_space - [inner_product_space ℝ E] [finite_dimensional ℝ E] [measurable_space E] [borel_space E] : + [normed_add_comm_group E] [inner_product_space ℝ E] [finite_dimensional ℝ E] + [measurable_space E] [borel_space E] : measure_space E := { volume := (std_orthonormal_basis ℝ E).to_basis.add_haar } diff --git a/src/measure_theory/measure/haar_of_inner.lean b/src/measure_theory/measure/haar_of_inner.lean index b10f73db4734b..0923810fae711 100644 --- a/src/measure_theory/measure/haar_of_inner.lean +++ b/src/measure_theory/measure/haar_of_inner.lean @@ -18,8 +18,9 @@ the canonical `volume` from the `measure_space` instance. open finite_dimensional measure_theory measure_theory.measure set -variables {ι F : Type*} [fintype ι] [inner_product_space ℝ F] [finite_dimensional ℝ F] -[measurable_space F] [borel_space F] +variables {ι F : Type*} +variables [fintype ι] [normed_add_comm_group F] [inner_product_space ℝ F] [finite_dimensional ℝ F] + [measurable_space F] [borel_space F] section From 72e87ce0bf26144ceb122d280a43aa10f1d20abb Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Fri, 24 Mar 2023 21:41:32 +0000 Subject: [PATCH 0685/1291] fix(tactic/ring_exp): `X^(2^3) = X^8` not `X^6` (#18645) Apparently `ring_exp` didn't correctly handle the exponentiation of coefficients, returning the product instead of the power as a coefficient. It still returned the right `expr`, so the faulty value is only used when checking that two terms are the same and so we can add their coefficients. In other words, this bug could only be triggered when `ring_exp` ends up adding `X^(a^b)` with `X^(a*b)` where `a` and `b` are both numerals. Thanks to @alainchmt for reporting the bug! --- src/tactic/ring_exp.lean | 3 ++- test/ring_exp.lean | 4 ++++ 2 files changed, 6 insertions(+), 1 deletion(-) diff --git a/src/tactic/ring_exp.lean b/src/tactic/ring_exp.lean index e39e2ff0db77e..8ee73b46356a9 100644 --- a/src/tactic/ring_exp.lean +++ b/src/tactic/ring_exp.lean @@ -1003,7 +1003,8 @@ meta def pow_coeff (p_p q_p : expr) (p q : coeff) : ring_exp_m (ex prod) := do ctx ← get_context, pq' ← mk_pow [p_p, q_p], (pq_p, pq_pf) ← lift $ norm_num.eval_pow pq', - pure $ ex.coeff ⟨pq_p, pq_p, pq_pf⟩ ⟨p.1 * q.1⟩ + if q.value.denom ≠ 1 then lift $ fail!"Only integer powers are supported, not {q.value}." + else pure $ ex.coeff ⟨pq_p, pq_p, pq_pf⟩ ⟨p.1 ^ q.value.num⟩ /-- Exponentiate two expressions. diff --git a/test/ring_exp.lean b/test/ring_exp.lean index 2485f20493a06..e2a9036b53ba7 100644 --- a/test/ring_exp.lean +++ b/test/ring_exp.lean @@ -79,6 +79,10 @@ by ring_exp_eq -- power does not have to be a syntactic match to `monoid.has_pow` example {α} [comm_ring α] (x : ℕ → α) : (x ^ 2 * x) = x ^ 3 := by ring_exp +-- Powers in the exponent get evaluated correctly. +example (X : ℤ) : (X^5 + 1) * (X^2^3 + X) = X^13 + X^8 + X^6 + X := +by ring_exp + end exponentiation section power_of_sum From cb3ceec8485239a61ed51d944cb9a95b68c6bafc Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Sat, 25 Mar 2023 06:31:08 +0000 Subject: [PATCH 0686/1291] chore(*): add mathlib4 synchronization comments (#18646) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.category.Group.basic` * `algebra.star.free` * `analysis.convex.segment` * `analysis.convex.star` * `category_theory.concrete_category.elementwise` * `category_theory.limits.concrete_category` * `category_theory.limits.shapes.diagonal` * `category_theory.limits.shapes.types` * `category_theory.localization.construction` * `category_theory.monad.basic` * `category_theory.monoidal.natural_transformation` * `category_theory.morphism_property` * `category_theory.sites.grothendieck` * `data.finset.sym` * `data.mv_polynomial.equiv` * `data.polynomial.field_division` * `group_theory.free_product` * `information_theory.hamming` * `linear_algebra.affine_space.affine_subspace` * `linear_algebra.affine_space.midpoint` * `linear_algebra.affine_space.restrict` * `number_theory.fermat4` * `number_theory.padics.padic_norm` * `ring_theory.polynomial.content` --- src/algebra/category/Group/basic.lean | 3 +++ src/algebra/star/free.lean | 3 +++ src/analysis/convex/segment.lean | 3 +++ src/analysis/convex/star.lean | 3 +++ src/category_theory/concrete_category/elementwise.lean | 3 +++ src/category_theory/limits/concrete_category.lean | 3 +++ src/category_theory/limits/shapes/diagonal.lean | 3 +++ src/category_theory/limits/shapes/types.lean | 3 +++ src/category_theory/localization/construction.lean | 3 +++ src/category_theory/monad/basic.lean | 3 +++ src/category_theory/monoidal/natural_transformation.lean | 3 +++ src/category_theory/morphism_property.lean | 3 +++ src/category_theory/sites/grothendieck.lean | 3 +++ src/data/finset/sym.lean | 3 +++ src/data/mv_polynomial/equiv.lean | 3 +++ src/data/polynomial/field_division.lean | 3 +++ src/group_theory/free_product.lean | 3 +++ src/information_theory/hamming.lean | 3 +++ src/linear_algebra/affine_space/affine_subspace.lean | 3 +++ src/linear_algebra/affine_space/midpoint.lean | 3 +++ src/linear_algebra/affine_space/restrict.lean | 3 +++ src/number_theory/fermat4.lean | 3 +++ src/number_theory/padics/padic_norm.lean | 3 +++ src/ring_theory/polynomial/content.lean | 3 +++ 24 files changed, 72 insertions(+) diff --git a/src/algebra/category/Group/basic.lean b/src/algebra/category/Group/basic.lean index 17a1980b15b3d..a03a9c64f8360 100644 --- a/src/algebra/category/Group/basic.lean +++ b/src/algebra/category/Group/basic.lean @@ -9,6 +9,9 @@ import category_theory.endomorphism /-! # Category instances for group, add_group, comm_group, and add_comm_group. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We introduce the bundled categories: * `Group` * `AddGroup` diff --git a/src/algebra/star/free.lean b/src/algebra/star/free.lean index 670cce3b058be..fca485ba14203 100644 --- a/src/algebra/star/free.lean +++ b/src/algebra/star/free.lean @@ -9,6 +9,9 @@ import algebra.free_algebra /-! # A *-algebra structure on the free algebra. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Reversing words gives a *-structure on the free monoid or on the free algebra on a type. ## Implementation note diff --git a/src/analysis/convex/segment.lean b/src/analysis/convex/segment.lean index 7eb8ab93d2ef1..6e0d57aed67d8 100644 --- a/src/analysis/convex/segment.lean +++ b/src/analysis/convex/segment.lean @@ -12,6 +12,9 @@ import tactic.positivity /-! # Segments in vector spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In a 𝕜-vector space, we define the following objects and properties. * `segment 𝕜 x y`: Closed segment joining `x` and `y`. * `open_segment 𝕜 x y`: Open segment joining `x` and `y`. diff --git a/src/analysis/convex/star.lean b/src/analysis/convex/star.lean index e0a9eaeb1ec41..6544563a61111 100644 --- a/src/analysis/convex/star.lean +++ b/src/analysis/convex/star.lean @@ -8,6 +8,9 @@ import analysis.convex.segment /-! # Star-convex sets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This files defines star-convex sets (aka star domains, star-shaped set, radially convex set). A set is star-convex at `x` if every segment from `x` to a point in the set is contained in the set. diff --git a/src/category_theory/concrete_category/elementwise.lean b/src/category_theory/concrete_category/elementwise.lean index 37b058d7edca5..e41d647a4eaed 100644 --- a/src/category_theory/concrete_category/elementwise.lean +++ b/src/category_theory/concrete_category/elementwise.lean @@ -10,6 +10,9 @@ import category_theory.concrete_category.basic import tactic.fresh_names /-! +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we provide various simp lemmas in its elementwise form via `tactic.elementwise`. -/ diff --git a/src/category_theory/limits/concrete_category.lean b/src/category_theory/limits/concrete_category.lean index 01d3a7451738a..4349caebfe5ba 100644 --- a/src/category_theory/limits/concrete_category.lean +++ b/src/category_theory/limits/concrete_category.lean @@ -13,6 +13,9 @@ import tactic.apply_fun /-! # Facts about (co)limits of functors into concrete categories + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ universes w v u diff --git a/src/category_theory/limits/shapes/diagonal.lean b/src/category_theory/limits/shapes/diagonal.lean index 40a2950038f3a..012e76ec3148c 100644 --- a/src/category_theory/limits/shapes/diagonal.lean +++ b/src/category_theory/limits/shapes/diagonal.lean @@ -10,6 +10,9 @@ import category_theory.limits.shapes.comm_sq /-! # The diagonal object of a morphism. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We provide various API and isomorphisms considering the diagonal object `Δ_{Y/X} := pullback f f` of a morphism `f : X ⟶ Y`. diff --git a/src/category_theory/limits/shapes/types.lean b/src/category_theory/limits/shapes/types.lean index c92488d57f25c..4666f03ce1a24 100644 --- a/src/category_theory/limits/shapes/types.lean +++ b/src/category_theory/limits/shapes/types.lean @@ -13,6 +13,9 @@ import tactic.elementwise /-! # Special shapes for limits in `Type`. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The general shape (co)limits defined in `category_theory.limits.types` are intended for use through the limits API, and the actual implementation should mostly be considered "sealed". diff --git a/src/category_theory/localization/construction.lean b/src/category_theory/localization/construction.lean index 6d2353242076c..e0d318df534db 100644 --- a/src/category_theory/localization/construction.lean +++ b/src/category_theory/localization/construction.lean @@ -11,6 +11,9 @@ import category_theory.category.Quiv # Construction of the localized category +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file constructs the localized category, obtained by formally inverting a class of maps `W : morphism_property C` in a category `C`. diff --git a/src/category_theory/monad/basic.lean b/src/category_theory/monad/basic.lean index 2516f941e3bc9..3ed6b2bcf7ec4 100644 --- a/src/category_theory/monad/basic.lean +++ b/src/category_theory/monad/basic.lean @@ -10,6 +10,9 @@ import category_theory.functor.reflects_isomorphisms /-! # Monads +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We construct the categories of monads and comonads, and their forgetful functors to endofunctors. (Note that these are the category theorist's monads, not the programmers monads. diff --git a/src/category_theory/monoidal/natural_transformation.lean b/src/category_theory/monoidal/natural_transformation.lean index d4a4776b6313b..478f462028274 100644 --- a/src/category_theory/monoidal/natural_transformation.lean +++ b/src/category_theory/monoidal/natural_transformation.lean @@ -9,6 +9,9 @@ import category_theory.full_subcategory /-! # Monoidal natural transformations +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Natural transformations between (lax) monoidal functors must satisfy an additional compatibility relation with the tensorators: `F.μ X Y ≫ app (X ⊗ Y) = (app X ⊗ app Y) ≫ G.μ X Y`. diff --git a/src/category_theory/morphism_property.lean b/src/category_theory/morphism_property.lean index a690812286cc4..bcf784ed4caf7 100644 --- a/src/category_theory/morphism_property.lean +++ b/src/category_theory/morphism_property.lean @@ -11,6 +11,9 @@ import category_theory.concrete_category.basic /-! # Properties of morphisms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We provide the basic framework for talking about properties of morphisms. The following meta-properties are defined diff --git a/src/category_theory/sites/grothendieck.lean b/src/category_theory/sites/grothendieck.lean index 658de4af3c30f..18277bfba4c8e 100644 --- a/src/category_theory/sites/grothendieck.lean +++ b/src/category_theory/sites/grothendieck.lean @@ -13,6 +13,9 @@ import order.copy /-! # Grothendieck topologies +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Definition and lemmas about Grothendieck topologies. A Grothendieck topology for a category `C` is a set of sieves on each object `X` satisfying certain closure conditions. diff --git a/src/data/finset/sym.lean b/src/data/finset/sym.lean index ec506a8667d3d..9b6246c51e858 100644 --- a/src/data/finset/sym.lean +++ b/src/data/finset/sym.lean @@ -11,6 +11,9 @@ import data.sym.sym2 /-! # Symmetric powers of a finset +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the symmetric powers of a finset as `finset (sym α n)` and `finset (sym2 α)`. ## Main declarations diff --git a/src/data/mv_polynomial/equiv.lean b/src/data/mv_polynomial/equiv.lean index a28475c1f0328..f4dadd5a7c302 100644 --- a/src/data/mv_polynomial/equiv.lean +++ b/src/data/mv_polynomial/equiv.lean @@ -15,6 +15,9 @@ import algebra.big_operators.fin /-! # Equivalences between polynomial rings +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file establishes a number of equivalences between polynomial rings, based on equivalences between the underlying types. diff --git a/src/data/polynomial/field_division.lean b/src/data/polynomial/field_division.lean index b1a6db3f07905..ea42eea04f17c 100644 --- a/src/data/polynomial/field_division.lean +++ b/src/data/polynomial/field_division.lean @@ -10,6 +10,9 @@ import ring_theory.euclidean_domain /-! # Theory of univariate polynomials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file starts looking like the ring theory of $ R[X] $ -/ diff --git a/src/group_theory/free_product.lean b/src/group_theory/free_product.lean index d7e1d926d1d2f..9426a82491c56 100644 --- a/src/group_theory/free_product.lean +++ b/src/group_theory/free_product.lean @@ -13,6 +13,9 @@ import data.set.pointwise.smul /-! # The free product of groups or monoids +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given an `ι`-indexed family `M` of monoids, we define their free product (categorical coproduct) `free_product M`. When `ι` and all `M i` have decidable equality, the free product bijects with the type `word M` of reduced words. This bijection is constructed by defining an action of diff --git a/src/information_theory/hamming.lean b/src/information_theory/hamming.lean index c1aacfc128e44..dcda8324f0b19 100644 --- a/src/information_theory/hamming.lean +++ b/src/information_theory/hamming.lean @@ -9,6 +9,9 @@ import analysis.normed.group.basic /-! # Hamming spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The Hamming metric counts the number of places two members of a (finite) Pi type differ. The Hamming norm is the same as the Hamming metric over additive groups, and counts the number of places a member of a (finite) Pi type differs from zero. diff --git a/src/linear_algebra/affine_space/affine_subspace.lean b/src/linear_algebra/affine_space/affine_subspace.lean index 6649e5d2a692f..321ca954d5a12 100644 --- a/src/linear_algebra/affine_space/affine_subspace.lean +++ b/src/linear_algebra/affine_space/affine_subspace.lean @@ -8,6 +8,9 @@ import linear_algebra.affine_space.affine_equiv /-! # Affine spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines affine subspaces (over modules) and the affine span of a set of points. ## Main definitions diff --git a/src/linear_algebra/affine_space/midpoint.lean b/src/linear_algebra/affine_space/midpoint.lean index 727d674d2b8e0..686bc7ebb701a 100644 --- a/src/linear_algebra/affine_space/midpoint.lean +++ b/src/linear_algebra/affine_space/midpoint.lean @@ -9,6 +9,9 @@ import linear_algebra.affine_space.affine_equiv /-! # Midpoint of a segment +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions * `midpoint R x y`: midpoint of the segment `[x, y]`. We define it for `x` and `y` diff --git a/src/linear_algebra/affine_space/restrict.lean b/src/linear_algebra/affine_space/restrict.lean index 5a82ba53e4136..b115879676dcf 100644 --- a/src/linear_algebra/affine_space/restrict.lean +++ b/src/linear_algebra/affine_space/restrict.lean @@ -8,6 +8,9 @@ import linear_algebra.affine_space.affine_subspace /-! # Affine map restrictions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines restrictions of affine maps. ## Main definitions diff --git a/src/number_theory/fermat4.lean b/src/number_theory/fermat4.lean index 38e16e2234eec..e802a60fc2046 100644 --- a/src/number_theory/fermat4.lean +++ b/src/number_theory/fermat4.lean @@ -9,6 +9,9 @@ import tactic.linear_combination /-! # Fermat's Last Theorem for the case n = 4 + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. There are no non-zero integers `a`, `b` and `c` such that `a ^ 4 + b ^ 4 = c ^ 4`. -/ diff --git a/src/number_theory/padics/padic_norm.lean b/src/number_theory/padics/padic_norm.lean index 4629482bd257e..6b5c5cd89fc68 100644 --- a/src/number_theory/padics/padic_norm.lean +++ b/src/number_theory/padics/padic_norm.lean @@ -8,6 +8,9 @@ import number_theory.padics.padic_val /-! # p-adic norm +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the `p`-adic norm on `ℚ`. The `p`-adic valuation on `ℚ` is the difference of the multiplicities of `p` in the numerator and diff --git a/src/ring_theory/polynomial/content.lean b/src/ring_theory/polynomial/content.lean index 89250e33b1380..121e7934aa498 100644 --- a/src/ring_theory/polynomial/content.lean +++ b/src/ring_theory/polynomial/content.lean @@ -11,6 +11,9 @@ import data.polynomial.cancel_leads /-! # GCD structures on polynomials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Definitions and basic results about polynomials over GCD domains, particularly their contents and primitive polynomials. From 0ca15f462f977d5385cb1d3b15b106f8adb6ff5f Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 25 Mar 2023 10:14:26 +0000 Subject: [PATCH 0687/1291] chore(category_theory/limits/shapes/concrete_category): Delete empty file (#18653) This file was made empty in #9864 but not deleted. --- .../limits/shapes/concrete_category.lean | 20 ------------------- 1 file changed, 20 deletions(-) delete mode 100644 src/category_theory/limits/shapes/concrete_category.lean diff --git a/src/category_theory/limits/shapes/concrete_category.lean b/src/category_theory/limits/shapes/concrete_category.lean deleted file mode 100644 index 9431a5aa1e86f..0000000000000 --- a/src/category_theory/limits/shapes/concrete_category.lean +++ /dev/null @@ -1,20 +0,0 @@ -/- -Copyright (c) 2017 Scott Morrison. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Scott Morrison --/ -import category_theory.limits.shapes.kernels -import category_theory.concrete_category.basic -import category_theory.concrete_category.elementwise - -/-! -# Facts about limits of functors into concrete categories - -This file doesn't yet attempt to be exhaustive; -it just contains lemmas that are useful -while comparing categorical limits with existing constructions in concrete categories. --/ - -universes u - -open category_theory From 517cc149e0b515d2893baa376226ed10feb319c7 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 25 Mar 2023 12:39:05 +0000 Subject: [PATCH 0688/1291] =?UTF-8?q?feat(data/finset/pointwise):=20`s=20?= =?UTF-8?q?=E2=88=A9=20t=20*=20s=20=E2=88=AA=20t=20=E2=8A=86=20s=20*=20t`?= =?UTF-8?q?=20(#17961)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit and distributivity of `set.to_finset`/`set.finite.to_finset` over algebraic operations. --- src/data/finset/n_ary.lean | 32 +++++++- src/data/finset/pointwise.lean | 100 ++++++++++++++++++++++++- src/data/polynomial/ring_division.lean | 6 +- src/data/set/finite.lean | 1 - src/data/set/n_ary.lean | 13 ++++ src/data/set/pointwise/basic.lean | 13 +++- src/data/set/pointwise/finite.lean | 6 ++ src/data/set/pointwise/smul.lean | 9 +++ 8 files changed, 170 insertions(+), 10 deletions(-) diff --git a/src/data/finset/n_ary.lean b/src/data/finset/n_ary.lean index 80aa9c093ba97..80345cf7443ca 100644 --- a/src/data/finset/n_ary.lean +++ b/src/data/finset/n_ary.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ import data.finset.prod +import data.set.finite /-! # N-ary images of finsets @@ -25,10 +26,10 @@ and `set.image2` already fulfills this task. open function set -namespace finset - variables {α α' β β' γ γ' δ δ' ε ε' ζ ζ' ν : Type*} - [decidable_eq α'] [decidable_eq β'] [decidable_eq γ] [decidable_eq γ'] [decidable_eq δ] + +namespace finset +variables [decidable_eq α'] [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 : γ} @@ -343,4 +344,29 @@ lemma image₂_right_identity {f : γ → β → γ} {b : β} (h : ∀ a, f a b image₂ f s {b} = s := by rw [image₂_singleton_right, funext h, image_id'] +variables [decidable_eq α] [decidable_eq β] + +lemma image₂_inter_union_subset {f : α → α → β} {s t : finset α} (hf : ∀ a b, f a b = f b a) : + image₂ f (s ∩ t) (s ∪ t) ⊆ image₂ f s t := +coe_subset.1 $ by { push_cast, exact image2_inter_union_subset hf } + +lemma image₂_union_inter_subset {f : α → α → β} {s t : finset α} (hf : ∀ a b, f a b = f b a) : + image₂ f (s ∪ t) (s ∩ t) ⊆ image₂ f s t := +coe_subset.1 $ by { push_cast, exact image2_union_inter_subset hf } + end finset + +namespace set +variables [decidable_eq γ] {s : set α} {t : set β} + +@[simp] lemma to_finset_image2 (f : α → β → γ) (s : set α) (t : set β) [fintype s] [fintype t] + [fintype (image2 f s t)] : + (image2 f s t).to_finset = finset.image₂ f s.to_finset t.to_finset := +finset.coe_injective $ by simp + +lemma finite.to_finset_image2 (f : α → β → γ) (hs : s.finite) (ht : t.finite) + (hf := hs.image2 f ht) : + hf.to_finset = finset.image₂ f hs.to_finset ht.to_finset := +finset.coe_injective $ by simp + +end set diff --git a/src/data/finset/pointwise.lean b/src/data/finset/pointwise.lean index 1e750480d0b8e..ac4a45def0c2b 100644 --- a/src/data/finset/pointwise.lean +++ b/src/data/finset/pointwise.lean @@ -5,6 +5,7 @@ Authors: Floris van Doorn, Yaël Dillies -/ import data.finset.n_ary import data.finset.preimage +import data.set.pointwise.finite import data.set.pointwise.smul import data.set.pointwise.list_of_fn @@ -83,6 +84,7 @@ image_singleton _ _ @[to_additive] lemma subset_one_iff_eq : s ⊆ 1 ↔ s = ∅ ∨ s = 1 := subset_singleton_iff @[to_additive] lemma nonempty.subset_one_iff (h : s.nonempty) : s ⊆ 1 ↔ s = 1 := h.subset_singleton_iff +@[simp, to_additive] lemma card_one : (1 : finset α).card = 1 := card_singleton _ /-- The singleton operation as a `one_hom`. -/ @[to_additive "The singleton operation as a `zero_hom`."] @@ -336,10 +338,20 @@ localized "attribute [instance] finset.has_nsmul finset.has_npow finset.has_zsmu protected def semigroup [semigroup α] : semigroup (finset α) := coe_injective.semigroup _ coe_mul +section comm_semigroup +variables [comm_semigroup α] {s t : finset α} + /-- `finset α` is a `comm_semigroup` under pointwise operations if `α` is. -/ @[to_additive "`finset α` is an `add_comm_semigroup` under pointwise operations if `α` is. "] -protected def comm_semigroup [comm_semigroup α] : comm_semigroup (finset α) := -coe_injective.comm_semigroup _ coe_mul +protected def comm_semigroup : comm_semigroup (finset α) := coe_injective.comm_semigroup _ coe_mul + +@[to_additive] lemma inter_mul_union_subset : s ∩ t * (s ∪ t) ⊆ s * t := +image₂_inter_union_subset mul_comm + +@[to_additive] lemma union_mul_inter_subset : (s ∪ t) * (s ∩ t) ⊆ s * t := +image₂_union_inter_subset mul_comm + +end comm_semigroup section mul_one_class variables [mul_one_class α] @@ -959,6 +971,16 @@ end open_locale pointwise +@[to_additive] lemma image_smul_comm [decidable_eq β] [decidable_eq γ] [has_smul α β] [has_smul α γ] + (f : β → γ) (a : α) (s : finset β) : + (∀ b, f (a • b) = a • f b) → (a • s).image f = a • s.image f := +image_comm + +@[to_additive] lemma image_smul_distrib [decidable_eq α] [decidable_eq β] [monoid α] [monoid β] + [monoid_hom_class F α β] (f : F) (a : α) (s : finset α) : + (a • s).image f = f a • s.image f := +image_comm $ map_mul _ _ + section group variables [decidable_eq β] [group α] [mul_action α β] {s t : finset β} {a : α} {b : β} @@ -980,6 +1002,9 @@ by { simp_rw ←coe_subset, push_cast, exact set.set_smul_subset_iff } @[to_additive] lemma subset_smul_finset_iff : s ⊆ a • t ↔ a⁻¹ • s ⊆ t := by { simp_rw ←coe_subset, push_cast, exact set.subset_set_smul_iff } +@[simp, to_additive] lemma card_smul_finset (a : α) (s : finset β) : (a • s).card = s.card := +card_image_of_injective _ $ mul_action.injective _ + end group section group_with_zero @@ -1073,3 +1098,74 @@ by { simp_rw ←image_neg, exact image₂_image_left_comm neg_smul } end ring end finset + +open_locale pointwise + +namespace set +section has_one +variables [has_one α] + +@[simp, to_additive] lemma to_finset_one : (1 : set α).to_finset = 1 := rfl + +@[simp, to_additive] +lemma finite.to_finset_one (h : (1 : set α).finite := finite_one) : h.to_finset = 1 := +finite.to_finset_singleton _ + +end has_one + +section has_mul +variables [decidable_eq α] [has_mul α] {s t : set α} + +@[simp, to_additive] lemma to_finset_mul (s t : set α) [fintype s] [fintype t] [fintype ↥(s * t)] : + (s * t).to_finset = s.to_finset * t.to_finset := +to_finset_image2 _ _ _ + +@[to_additive] lemma finite.to_finset_mul (hs : s.finite) (ht : t.finite) (hf := hs.mul ht) : + hf.to_finset = hs.to_finset * ht.to_finset := +finite.to_finset_image2 _ _ _ + +end has_mul + +section has_smul +variables [has_smul α β] [decidable_eq β] {a : α} {s : set α} {t : set β} + +@[simp, to_additive] +lemma to_finset_smul (s : set α) (t : set β) [fintype s] [fintype t] [fintype ↥(s • t)] : + (s • t).to_finset = s.to_finset • t.to_finset := +to_finset_image2 _ _ _ + +@[to_additive] lemma finite.to_finset_smul (hs : s.finite) (ht : t.finite) (hf := hs.smul ht) : + hf.to_finset = hs.to_finset • ht.to_finset := +finite.to_finset_image2 _ _ _ + +end has_smul + +section has_smul +variables [decidable_eq β] [has_smul α β] {a : α} {s : set β} + +@[simp, to_additive] +lemma to_finset_smul_set (a : α) (s : set β) [fintype s] [fintype ↥(a • s)] : + (a • s).to_finset = a • s.to_finset := +to_finset_image _ _ + +@[to_additive] +lemma finite.to_finset_smul_set (hs : s.finite) (hf : (a • s).finite := hs.smul_set) : + hf.to_finset = a • hs.to_finset := +finite.to_finset_image _ _ _ + +end has_smul + +section has_vsub +variables [decidable_eq α] [has_vsub α β] {s t : set β} +include α + +@[simp] lemma to_finset_vsub (s t : set β) [fintype s] [fintype t] [fintype ↥(s -ᵥ t)] : + (s -ᵥ t : set α).to_finset = s.to_finset -ᵥ t.to_finset := +to_finset_image2 _ _ _ + +lemma finite.to_finset_vsub (hs : s.finite) (ht : t.finite) (hf := hs.vsub ht) : + hf.to_finset = hs.to_finset -ᵥ ht.to_finset := +finite.to_finset_image2 _ _ _ + +end has_vsub +end set diff --git a/src/data/polynomial/ring_division.lean b/src/data/polynomial/ring_division.lean index f64a1c21d0d01..d833e027586a3 100644 --- a/src/data/polynomial/ring_division.lean +++ b/src/data/polynomial/ring_division.lean @@ -676,9 +676,9 @@ by simp only [empty_eq_zero, pow_zero, nth_roots, ← C_1, ← C_sub, roots_C] lemma card_nth_roots (n : ℕ) (a : R) : (nth_roots n a).card ≤ n := -if hn : n = 0 -then if h : (X : R[X]) ^ n - C a = 0 - then by simp only [nat.zero_le, nth_roots, roots, h, dif_pos rfl, empty_eq_zero, card_zero] +if hn : n = 0 then + if h : (X : R[X]) ^ n - C a = 0 then + by simp only [nat.zero_le, nth_roots, roots, h, dif_pos rfl, empty_eq_zero, multiset.card_zero] else with_bot.coe_le_coe.1 (le_trans (card_roots h) (by { rw [hn, pow_zero, ← C_1, ← ring_hom.map_sub ], exact degree_C_le })) diff --git a/src/data/set/finite.lean b/src/data/set/finite.lean index 649d62f1ec157..fe1ca45972561 100644 --- a/src/data/set/finite.lean +++ b/src/data/set/finite.lean @@ -196,7 +196,6 @@ by { ext, simp } @[simp] protected lemma to_finset_empty (h : (∅ : set α).finite) : h.to_finset = ∅ := by { ext, simp } --- Note: Not `simp` because `set.finite.to_finset_set_of` already proves it @[simp] protected lemma to_finset_univ [fintype α] (h : (set.univ : set α).finite) : h.to_finset = finset.univ := by { ext, simp } diff --git a/src/data/set/n_ary.lean b/src/data/set/n_ary.lean index 6b0413ebb416a..27c35f92b03af 100644 --- a/src/data/set/n_ary.lean +++ b/src/data/set/n_ary.lean @@ -339,4 +339,17 @@ lemma image2_right_identity {f : α → β → α} {b : β} (h : ∀ a, f a b = image2 f s {b} = s := by rw [image2_singleton_right, funext h, image_id'] +lemma image2_inter_union_subset {f : α → α → β} {s t : set α} (hf : ∀ a b, f a b = f b a) : + image2 f (s ∩ t) (s ∪ t) ⊆ image2 f s t := +begin + rintro _ ⟨a, b, ha, hb | hb, rfl⟩, + { rw hf, + exact mem_image2_of_mem hb ha.2 }, + { exact mem_image2_of_mem ha.1 hb } +end + +lemma image2_union_inter_subset {f : α → α → β} {s t : set α} (hf : ∀ a b, f a b = f b a) : + image2 f (s ∪ t) (s ∩ t) ⊆ image2 f s t := +by { rw image2_comm hf, exact image2_inter_union_subset hf } + end set diff --git a/src/data/set/pointwise/basic.lean b/src/data/set/pointwise/basic.lean index 9cd15e343d3fe..aa0c9c0ac8a2f 100644 --- a/src/data/set/pointwise/basic.lean +++ b/src/data/set/pointwise/basic.lean @@ -389,12 +389,23 @@ protected def semigroup [semigroup α] : semigroup (set α) := { mul_assoc := λ _ _ _, image2_assoc mul_assoc, ..set.has_mul } +section comm_semigroup +variables [comm_semigroup α] {s t : set α} + /-- `set α` is a `comm_semigroup` under pointwise operations if `α` is. -/ @[to_additive "`set α` is an `add_comm_semigroup` under pointwise operations if `α` is."] -protected def comm_semigroup [comm_semigroup α] : comm_semigroup (set α) := +protected def comm_semigroup : comm_semigroup (set α) := { mul_comm := λ s t, image2_comm mul_comm ..set.semigroup } +@[to_additive] lemma inter_mul_union_subset : (s ∩ t) * (s ∪ t) ⊆ s * t := +image2_inter_union_subset mul_comm + +@[to_additive] lemma union_mul_inter_subset : (s ∪ t) * (s ∩ t) ⊆ s * t := +image2_union_inter_subset mul_comm + +end comm_semigroup + section mul_one_class variables [mul_one_class α] diff --git a/src/data/set/pointwise/finite.lean b/src/data/set/pointwise/finite.lean index cb402a79ba3ed..741a5f60cf39e 100644 --- a/src/data/set/pointwise/finite.lean +++ b/src/data/set/pointwise/finite.lean @@ -16,6 +16,12 @@ open_locale pointwise variables {F α β γ : Type*} namespace set +section has_one +variables [has_one α] + +@[simp, to_additive] lemma finite_one : (1 : set α).finite := finite_singleton _ + +end has_one section has_involutive_inv variables [has_involutive_inv α] {s : set α} diff --git a/src/data/set/pointwise/smul.lean b/src/data/set/pointwise/smul.lean index a4a842ee19225..0357b1676fd5b 100644 --- a/src/data/set/pointwise/smul.lean +++ b/src/data/set/pointwise/smul.lean @@ -378,6 +378,15 @@ end vsub open_locale pointwise +@[to_additive] lemma image_smul_comm [has_smul α β] [has_smul α γ] (f : β → γ) (a : α) (s : set β) : + (∀ b, f (a • b) = a • f b) → f '' (a • s) = a • f '' s := +image_comm + +@[to_additive] lemma image_smul_distrib [mul_one_class α] [mul_one_class β] [monoid_hom_class F α β] + (f : F) (a : α) (s : set α) : + f '' (a • s) = f a • f '' s := +image_comm $ map_mul _ _ + section smul_with_zero variables [has_zero α] [has_zero β] [smul_with_zero α β] {s : set α} {t : set β} From 0e3aacdc98d25e0afe035c452d876d28cbffaa7e Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 25 Mar 2023 15:44:09 +0000 Subject: [PATCH 0689/1291] feat(analysis/convex/topology): Closure of an open segment (#18589) The closure of an open segment is the corresponding closed segment. --- src/analysis/convex/topology.lean | 33 ++++++++++++++++++++++++++++--- 1 file changed, 30 insertions(+), 3 deletions(-) diff --git a/src/analysis/convex/topology.lean b/src/analysis/convex/topology.lean index f36f01234f2f7..cc9d7665e1d2a 100644 --- a/src/analysis/convex/topology.lean +++ b/src/analysis/convex/topology.lean @@ -22,11 +22,11 @@ We prove the following facts: assert_not_exists has_norm -variables {ι : Type*} {E : Type*} - open metric set open_locale pointwise convex +variables {ι 𝕜 E : Type*} + lemma real.convex_iff_is_preconnected {s : set ℝ} : convex ℝ s ↔ is_preconnected s := convex_iff_ord_connected.trans is_preconnected_iff_ord_connected.symm @@ -69,9 +69,36 @@ end std_simplex /-! ### Topological vector space -/ +section topological_space +variables [linear_ordered_ring 𝕜] [densely_ordered 𝕜] [topological_space 𝕜] [order_topology 𝕜] + [add_comm_group E] [topological_space E] [has_continuous_add E] [module 𝕜 E] + [has_continuous_smul 𝕜 E] {x y : E} + +lemma segment_subset_closure_open_segment : [x -[𝕜] y] ⊆ closure (open_segment 𝕜 x y) := +begin + rw [segment_eq_image, open_segment_eq_image, ←closure_Ioo (zero_ne_one' 𝕜)], + exact image_closure_subset_closure_image (by continuity), +end + +end topological_space + +section pseudo_metric_space +variables [linear_ordered_ring 𝕜] [densely_ordered 𝕜] [pseudo_metric_space 𝕜] [order_topology 𝕜] + [proper_space 𝕜] [compact_Icc_space 𝕜] [add_comm_group E] [topological_space E] [t2_space E] + [has_continuous_add E] [module 𝕜 E] [has_continuous_smul 𝕜 E] + +@[simp] lemma closure_open_segment (x y : E) : closure (open_segment 𝕜 x y) = [x -[𝕜] y] := +begin + rw [segment_eq_image, open_segment_eq_image, ←closure_Ioo (zero_ne_one' 𝕜)], + exact (image_closure_of_is_compact (bounded_Ioo _ _).is_compact_closure $ + continuous.continuous_on $ by continuity).symm, +end + +end pseudo_metric_space + section has_continuous_const_smul -variables {𝕜 : Type*} [linear_ordered_field 𝕜] [add_comm_group E] [module 𝕜 E] [topological_space E] +variables [linear_ordered_field 𝕜] [add_comm_group E] [module 𝕜 E] [topological_space E] [topological_add_group E] [has_continuous_const_smul 𝕜 E] /-- If `s` is a convex set, then `a • interior s + b • closure s ⊆ interior s` for all `0 < a`, From d3af0609f6db8691dffdc3e1fb7feb7da72698f2 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sat, 25 Mar 2023 15:44:10 +0000 Subject: [PATCH 0690/1291] chore(analysis/normed_space/basic): introduce `norm_smul_le` (#18650) Currently `norm_smul_le x y` is just a special case of `(norm_smul x y).le`; but if in future we generalize `normed_space` to work over normed rings, it will continue to hold where `norm_smul` no longer does. This adjusts downstream proofs to use the weaker lemma too when the stronger one isn't needed, both so that we have less to fix if/when we make the suggested refactor, and because the new spelling is shorter. This adds the corresponding `nnnorm` and `dist` and `nndist` lemmas too. --- src/analysis/normed_space/basic.lean | 55 +++++++++++++------ src/analysis/normed_space/completion.lean | 2 +- src/analysis/normed_space/operator_norm.lean | 4 +- .../normed_space/star/multiplier.lean | 2 +- src/analysis/quaternion.lean | 2 +- .../function/strongly_measurable/basic.lean | 2 +- src/topology/continuous_function/bounded.lean | 2 +- src/topology/continuous_function/compact.lean | 2 +- .../continuous_function/zero_at_infty.lean | 2 +- 9 files changed, 46 insertions(+), 27 deletions(-) diff --git a/src/analysis/normed_space/basic.lean b/src/analysis/normed_space/basic.lean index 9955d91785f4d..ee8c4968a9851 100644 --- a/src/analysis/normed_space/basic.lean +++ b/src/analysis/normed_space/basic.lean @@ -42,28 +42,47 @@ end prio variables [normed_field α] [seminormed_add_comm_group β] +-- note: while these are currently strictly weaker than the versions without `le`, they will cease +-- to be if we eventually generalize `normed_space` from `normed_field α` to `normed_ring α`. +section le + +lemma norm_smul_le [normed_space α β] (r : α) (x : β) : ‖r • x‖ ≤ ‖r‖ * ‖x‖ := +normed_space.norm_smul_le _ _ + +lemma nnnorm_smul_le [normed_space α β] (s : α) (x : β) : ‖s • x‖₊ ≤ ‖s‖₊ * ‖x‖₊ := +norm_smul_le s x + +lemma dist_smul_le [normed_space α β] (s : α) (x y : β) : dist (s • x) (s • y) ≤ ‖s‖ * dist x y := +by simpa only [dist_eq_norm, ←smul_sub] using norm_smul_le _ _ + +lemma nndist_smul_le [normed_space α β] (s : α) (x y : β) : + nndist (s • x) (s • y) ≤ ‖s‖₊ * nndist x y := +dist_smul_le s x y + +end le + @[priority 100] -- see Note [lower instance priority] instance normed_space.has_bounded_smul [normed_space α β] : has_bounded_smul α β := { dist_smul_pair' := λ x y₁ y₂, - by simpa [dist_eq_norm, smul_sub] using normed_space.norm_smul_le x (y₁ - y₂), + by simpa [dist_eq_norm, smul_sub] using norm_smul_le x (y₁ - y₂), dist_pair_smul' := λ x₁ x₂ y, - by simpa [dist_eq_norm, sub_smul] using normed_space.norm_smul_le (x₁ - x₂) y } + by simpa [dist_eq_norm, sub_smul] using norm_smul_le (x₁ - x₂) y } -- Shortcut instance, as otherwise this will be found by `normed_space.to_module` and be -- noncomputable. instance : module ℝ ℝ := by apply_instance instance normed_field.to_normed_space : normed_space α α := -{ norm_smul_le := λ a b, le_of_eq (norm_mul a b) } +{ norm_smul_le := λ a b, norm_mul_le a b } lemma norm_smul [normed_space α β] (s : α) (x : β) : ‖s • x‖ = ‖s‖ * ‖x‖ := begin by_cases h : s = 0, { simp [h] }, - { refine le_antisymm (normed_space.norm_smul_le s x) _, + { refine le_antisymm (norm_smul_le s x) _, calc ‖s‖ * ‖x‖ = ‖s‖ * ‖s⁻¹ • s • x‖ : by rw [inv_smul_smul₀ h] ... ≤ ‖s‖ * (‖s⁻¹‖ * ‖s • x‖) : - mul_le_mul_of_nonneg_left (normed_space.norm_smul_le _ _) (norm_nonneg _) + mul_le_mul_of_nonneg_left (norm_smul_le _ _) (norm_nonneg _) ... = ‖s • x‖ : by rw [norm_inv, ← mul_assoc, mul_inv_cancel (mt norm_eq_zero.1 h), one_mul] } end @@ -108,12 +127,12 @@ this.eventually (gt_mem_nhds h) lemma filter.tendsto.zero_smul_is_bounded_under_le {f : ι → α} {g : ι → E} {l : filter ι} (hf : tendsto f l (𝓝 0)) (hg : is_bounded_under (≤) l (norm ∘ g)) : tendsto (λ x, f x • g x) l (𝓝 0) := -hf.op_zero_is_bounded_under_le hg (•) (λ x y, (norm_smul x y).le) +hf.op_zero_is_bounded_under_le hg (•) norm_smul_le lemma filter.is_bounded_under.smul_tendsto_zero {f : ι → α} {g : ι → E} {l : filter ι} (hf : is_bounded_under (≤) l (norm ∘ f)) (hg : tendsto g l (𝓝 0)) : tendsto (λ x, f x • g x) l (𝓝 0) := -hg.op_zero_is_bounded_under_le hf (flip (•)) (λ x y, ((norm_smul y x).trans (mul_comm _ _)).le) +hg.op_zero_is_bounded_under_le hf (flip (•)) (λ x y, (norm_smul_le y x).trans_eq (mul_comm _ _)) theorem closure_ball [normed_space ℝ E] (x : E) {r : ℝ} (hr : r ≠ 0) : closure (ball x r) = closed_ball x r := @@ -230,26 +249,26 @@ by simp [homeomorph_unit_ball] open normed_field instance : normed_space α (ulift E) := -{ norm_smul_le := λ s x, (normed_space.norm_smul_le s x.down : _), +{ norm_smul_le := λ s x, (norm_smul_le s x.down : _), ..ulift.normed_add_comm_group, ..ulift.module' } /-- The product of two normed spaces is a normed space, with the sup norm. -/ instance prod.normed_space : normed_space α (E × F) := -{ norm_smul_le := λ s x, le_of_eq $ by simp [prod.norm_def, norm_smul, mul_max_of_nonneg], +{ norm_smul_le := λ s x, by simp [prod.norm_def, norm_smul_le, mul_max_of_nonneg], ..prod.normed_add_comm_group, ..prod.module } /-- The product of finitely many normed spaces is a normed space, with the sup norm. -/ instance pi.normed_space {E : ι → Type*} [fintype ι] [∀i, seminormed_add_comm_group (E i)] [∀i, normed_space α (E i)] : normed_space α (Πi, E i) := -{ norm_smul_le := λ a f, le_of_eq $ - show (↑(finset.sup finset.univ (λ (b : ι), ‖a • f b‖₊)) : ℝ) = - ‖a‖₊ * ↑(finset.sup finset.univ (λ (b : ι), ‖f b‖₊)), - by simp only [(nnreal.coe_mul _ _).symm, nnreal.mul_finset_sup, nnnorm_smul] } +{ norm_smul_le := λ a f, begin + simp_rw [←coe_nnnorm, ←nnreal.coe_mul, nnreal.coe_le_coe, pi.nnnorm_def, nnreal.mul_finset_sup], + exact finset.sup_mono_fun (λ _ _, norm_smul_le _ _), + end } instance mul_opposite.normed_space : normed_space α Eᵐᵒᵖ := -{ norm_smul_le := λ s x, (norm_smul s x.unop).le, +{ norm_smul_le := λ s x, norm_smul_le s x.unop, ..mul_opposite.normed_add_comm_group, ..mul_opposite.module _ } @@ -258,7 +277,7 @@ instance submodule.normed_space {𝕜 R : Type*} [has_smul 𝕜 R] [normed_field {E : Type*} [seminormed_add_comm_group E] [normed_space 𝕜 E] [module R E] [is_scalar_tower 𝕜 R E] (s : submodule R E) : normed_space 𝕜 s := -{ norm_smul_le := λc x, le_of_eq $ norm_smul c (x : E) } +{ norm_smul_le := λc x, norm_smul_le c (x : E) } /-- If there is a scalar `c` with `‖c‖>1`, then any element with nonzero norm can be moved by scalar multiplication to any shell of width `‖c‖`. Also recap information on the norm of @@ -304,7 +323,7 @@ See note [reducible non-instances] -/ def normed_space.induced {F : Type*} (α β γ : Type*) [normed_field α] [add_comm_group β] [module α β] [seminormed_add_comm_group γ] [normed_space α γ] [linear_map_class F α β γ] (f : F) : @normed_space α β _ (seminormed_add_comm_group.induced β γ f) := -{ norm_smul_le := λ a b, by {unfold norm, exact (map_smul f a b).symm ▸ (norm_smul a (f b)).le } } +{ norm_smul_le := λ a b, by {unfold norm, exact (map_smul f a b).symm ▸ norm_smul_le a (f b) } } section normed_add_comm_group @@ -549,7 +568,7 @@ See note [reducible non-instances] -/ def normed_algebra.induced {F : Type*} (α β γ : Type*) [normed_field α] [ring β] [algebra α β] [semi_normed_ring γ] [normed_algebra α γ] [non_unital_alg_hom_class F α β γ] (f : F) : @normed_algebra α β _ (semi_normed_ring.induced β γ f) := -{ norm_smul_le := λ a b, by {unfold norm, exact (map_smul f a b).symm ▸ (norm_smul a (f b)).le } } +{ norm_smul_le := λ a b, by {unfold norm, exact (map_smul f a b).symm ▸ norm_smul_le a (f b) } } instance subalgebra.to_normed_algebra {𝕜 A : Type*} [semi_normed_ring A] [normed_field 𝕜] [normed_algebra 𝕜 A] (S : subalgebra 𝕜 A) : normed_algebra 𝕜 S := @@ -569,7 +588,7 @@ instance {𝕜 : Type*} {𝕜' : Type*} {E : Type*} [I : normed_add_comm_group E /-- If `E` is a normed space over `𝕜'` and `𝕜` is a normed algebra over `𝕜'`, then `restrict_scalars.module` is additionally a `normed_space`. -/ instance : normed_space 𝕜 (restrict_scalars 𝕜 𝕜' E) := -{ norm_smul_le := λ c x, (normed_space.norm_smul_le (algebra_map 𝕜 𝕜' c) (_ : E)).trans_eq $ +{ norm_smul_le := λ c x, (norm_smul_le (algebra_map 𝕜 𝕜' c) (_ : E)).trans_eq $ by rw norm_algebra_map', ..restrict_scalars.module 𝕜 𝕜' E } diff --git a/src/analysis/normed_space/completion.lean b/src/analysis/normed_space/completion.lean index 11b159a3fcf99..bac77e02d1122 100644 --- a/src/analysis/normed_space/completion.lean +++ b/src/analysis/normed_space/completion.lean @@ -91,7 +91,7 @@ instance [semi_normed_comm_ring A] [normed_algebra 𝕜 A] [has_uniform_continuo { exact is_closed_le (continuous.comp (continuous_norm) (continuous_const_smul r)) (continuous.comp (continuous_mul_left _) continuous_norm), }, { intros x, - simp only [← coe_smul, norm_coe], exact normed_space.norm_smul_le r x } + simp only [← coe_smul, norm_coe], exact norm_smul_le r x } end, ..completion.algebra A 𝕜} diff --git a/src/analysis/normed_space/operator_norm.lean b/src/analysis/normed_space/operator_norm.lean index b710ee3a5e061..20b8d79b5e7d4 100644 --- a/src/analysis/normed_space/operator_norm.lean +++ b/src/analysis/normed_space/operator_norm.lean @@ -976,7 +976,7 @@ variables (𝕜) (𝕜' : Type*) [normed_field 𝕜'] [normed_algebra 𝕜 𝕜' /-- Scalar multiplication as a continuous bilinear map. -/ def lsmul : 𝕜' →L[𝕜] E →L[𝕜] E := ((algebra.lsmul 𝕜 E).to_linear_map : 𝕜' →ₗ[𝕜] E →ₗ[𝕜] E).mk_continuous₂ 1 $ - λ c x, by simpa only [one_mul] using (norm_smul c x).le + λ c x, by simpa only [one_mul] using norm_smul_le c x @[simp] lemma lsmul_apply (c : 𝕜') (x : E) : lsmul 𝕜 𝕜' c x = c • x := rfl @@ -994,7 +994,7 @@ end variables {𝕜} lemma op_norm_lsmul_apply_le (x : 𝕜') : ‖(lsmul 𝕜 𝕜' x : E →L[𝕜] E)‖ ≤ ‖x‖ := -continuous_linear_map.op_norm_le_bound _ (norm_nonneg x) $ λ y, (norm_smul x y).le +continuous_linear_map.op_norm_le_bound _ (norm_nonneg x) $ λ y, norm_smul_le x y /-- The norm of `lsmul` is at most 1 in any semi-normed group. -/ lemma op_norm_lsmul_le : ‖(lsmul 𝕜 𝕜' : 𝕜' →L[𝕜] E →L[𝕜] E)‖ ≤ 1 := diff --git a/src/analysis/normed_space/star/multiplier.lean b/src/analysis/normed_space/star/multiplier.lean index a72a923aeb97a..039c59151145b 100644 --- a/src/analysis/normed_space/star/multiplier.lean +++ b/src/analysis/normed_space/star/multiplier.lean @@ -383,7 +383,7 @@ lemma norm_def' (a : 𝓜(𝕜, A)) : ‖a‖ = ‖a.to_prod_mul_opposite_hom‖ lemma nnnorm_def' (a : 𝓜(𝕜, A)) : ‖a‖₊ = ‖a.to_prod_mul_opposite_hom‖₊ := rfl instance : normed_space 𝕜 𝓜(𝕜, A) := -{ norm_smul_le := λ k a, (norm_smul k a.to_prod_mul_opposite).le, +{ norm_smul_le := λ k a, norm_smul_le k a.to_prod_mul_opposite, .. double_centralizer.module } instance : normed_algebra 𝕜 𝓜(𝕜, A) := diff --git a/src/analysis/quaternion.lean b/src/analysis/quaternion.lean index 1fedf745aa37b..02fe08b71b90e 100644 --- a/src/analysis/quaternion.lean +++ b/src/analysis/quaternion.lean @@ -77,7 +77,7 @@ noncomputable instance : normed_division_ring ℍ := exact real.sqrt_mul norm_sq_nonneg _ } } instance : normed_algebra ℝ ℍ := -{ norm_smul_le := λ a x, (norm_smul a x).le, +{ norm_smul_le := norm_smul_le, to_algebra := (quaternion.algebra : algebra ℝ ℍ) } instance : cstar_ring ℍ := diff --git a/src/measure_theory/function/strongly_measurable/basic.lean b/src/measure_theory/function/strongly_measurable/basic.lean index 1e406c67fbf5f..425643486a0f9 100644 --- a/src/measure_theory/function/strongly_measurable/basic.lean +++ b/src/measure_theory/function/strongly_measurable/basic.lean @@ -234,7 +234,7 @@ lemma norm_approx_bounded_le {β} {f : α → β} [seminormed_add_comm_group β] ‖hf.approx_bounded c n x‖ ≤ c := begin simp only [strongly_measurable.approx_bounded, simple_func.coe_map, function.comp_app], - refine (norm_smul _ _).le.trans _, + refine (norm_smul_le _ _).trans _, by_cases h0 : ‖hf.approx n x‖ = 0, { simp only [h0, div_zero, min_eq_right, zero_le_one, norm_zero, mul_zero], exact hc, }, diff --git a/src/topology/continuous_function/bounded.lean b/src/topology/continuous_function/bounded.lean index da3985efeeb04..32146f3e3f16c 100644 --- a/src/topology/continuous_function/bounded.lean +++ b/src/topology/continuous_function/bounded.lean @@ -1165,7 +1165,7 @@ functions from `α` to `𝕜`. -/ instance has_smul' : has_smul (α →ᵇ 𝕜) (α →ᵇ β) := ⟨λ (f : α →ᵇ 𝕜) (g : α →ᵇ β), of_normed_add_comm_group (λ x, (f x) • (g x)) (f.continuous.smul g.continuous) (‖f‖ * ‖g‖) (λ x, calc - ‖f x • g x‖ ≤ ‖f x‖ * ‖g x‖ : normed_space.norm_smul_le _ _ + ‖f x • g x‖ ≤ ‖f x‖ * ‖g x‖ : norm_smul_le _ _ ... ≤ ‖f‖ * ‖g‖ : mul_le_mul (f.norm_coe_le_norm _) (g.norm_coe_le_norm _) (norm_nonneg _) (norm_nonneg _)) ⟩ diff --git a/src/topology/continuous_function/compact.lean b/src/topology/continuous_function/compact.lean index d6c6eff9e0520..7b83f26c73a94 100644 --- a/src/topology/continuous_function/compact.lean +++ b/src/topology/continuous_function/compact.lean @@ -227,7 +227,7 @@ section variables {𝕜 : Type*} [normed_field 𝕜] [normed_space 𝕜 E] instance : normed_space 𝕜 C(α,E) := -{ norm_smul_le := λ c f, le_of_eq (norm_smul c (mk_of_compact f)) } +{ norm_smul_le := λ c f, (norm_smul_le c (mk_of_compact f) : _) } section variables (α 𝕜 E) diff --git a/src/topology/continuous_function/zero_at_infty.lean b/src/topology/continuous_function/zero_at_infty.lean index aa34377be2ed4..2eca0c8ff9a92 100644 --- a/src/topology/continuous_function/zero_at_infty.lean +++ b/src/topology/continuous_function/zero_at_infty.lean @@ -411,7 +411,7 @@ normed_add_comm_group.induced C₀(α, β) (α →ᵇ β) (⟨to_bcf, rfl, λ x lemma norm_to_bcf_eq_norm {f : C₀(α, β)} : ‖f.to_bcf‖ = ‖f‖ := rfl instance : normed_space 𝕜 C₀(α, β) := -{ norm_smul_le := λ k f, (norm_smul k f.to_bcf).le } +{ norm_smul_le := λ k f, norm_smul_le k f.to_bcf } end normed_space From 55d771df074d0dd020139ee1cd4b95521422df9f Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sat, 25 Mar 2023 17:27:44 +0000 Subject: [PATCH 0691/1291] =?UTF-8?q?feat(topology/*):=20add=20lemmas=20ab?= =?UTF-8?q?out=20`=F0=9D=93=9D[=E2=8B=83=20i,=20s=20i]=20a`=20(#18321)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit * Add `theorem nhds_within_eq_nhds`, `nhds_within_bUnion`, `nhds_within_sUnion`, `nhds_within_Union`, `nhds_within_inter_of_mem'`. * Add `locally_finite.nhds_within_Union`, use it to golf `locally_finite.is_closed_Union` and `locally_finite.closure_Union`. * Reformulate `continuous_subtype_nhds_cover` in terms of `continuous_on`, rename to `continuous_of_cover_nhds`. * Reformulate `continuous_subtype_is_closed_cover` in terms of `continuous_on`, several versions are named `locally_finite.continuous_on_Union`, `locally_finite.continuous`, and primed versions of these lemmas. * Reorder imports. Co-authored-by: Yaël Dillies --- src/data/analysis/topology.lean | 1 + src/topology/constructions.lean | 38 ------------ src/topology/continuous_function/basic.lean | 8 +-- src/topology/continuous_on.lean | 27 +++++++- src/topology/locally_finite.lean | 69 +++++++++++++++------ src/topology/subset_properties.lean | 1 + 6 files changed, 82 insertions(+), 62 deletions(-) diff --git a/src/data/analysis/topology.lean b/src/data/analysis/topology.lean index c76fd7b2ac460..f66aac0b96661 100644 --- a/src/data/analysis/topology.lean +++ b/src/data/analysis/topology.lean @@ -5,6 +5,7 @@ Authors: Mario Carneiro -/ import data.analysis.filter import topology.bases +import topology.locally_finite /-! # Computational realization of topological spaces (experimental) diff --git a/src/topology/constructions.lean b/src/topology/constructions.lean index 3c8444fafec4e..c417015bb6d40 100644 --- a/src/topology/constructions.lean +++ b/src/topology/constructions.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro, Patrick Massot -/ import topology.maps -import topology.locally_finite import order.filter.pi /-! @@ -866,43 +865,6 @@ lemma tendsto_subtype_rng {β : Type*} {p : α → Prop} {b : filter β} {f : β ∀{a:subtype p}, tendsto f b (𝓝 a) ↔ tendsto (λx, (f x : α)) b (𝓝 (a : α)) | ⟨a, ha⟩ := by rw [nhds_subtype_eq_comap, tendsto_comap_iff, subtype.coe_mk] -lemma continuous_subtype_nhds_cover {ι : Sort*} {f : α → β} {c : ι → α → Prop} - (c_cover : ∀x:α, ∃i, {x | c i x} ∈ 𝓝 x) - (f_cont : ∀i, continuous (λ(x : subtype (c i)), f x)) : - continuous f := -continuous_iff_continuous_at.mpr $ assume x, - let ⟨i, (c_sets : {x | c i x} ∈ 𝓝 x)⟩ := c_cover x in - let x' : subtype (c i) := ⟨x, mem_of_mem_nhds c_sets⟩ in - calc map f (𝓝 x) = map f (map coe (𝓝 x')) : - congr_arg (map f) (map_nhds_subtype_coe_eq _ $ c_sets).symm - ... = map (λx:subtype (c i), f x) (𝓝 x') : rfl - ... ≤ 𝓝 (f x) : continuous_iff_continuous_at.mp (f_cont i) x' - -lemma continuous_subtype_is_closed_cover {ι : Sort*} {f : α → β} (c : ι → α → Prop) - (h_lf : locally_finite (λi, {x | c i x})) - (h_is_closed : ∀i, is_closed {x | c i x}) - (h_cover : ∀x, ∃i, c i x) - (f_cont : ∀i, continuous (λ(x : subtype (c i)), f x)) : - continuous f := -continuous_iff_is_closed.mpr $ - assume s hs, - have ∀i, is_closed ((coe : {x | c i x} → α) '' (f ∘ coe ⁻¹' s)), - from assume i, - (closed_embedding_subtype_coe (h_is_closed _)).is_closed_map _ (hs.preimage (f_cont i)), - have is_closed (⋃i, (coe : {x | c i x} → α) '' (f ∘ coe ⁻¹' s)), - from locally_finite.is_closed_Union - (h_lf.subset $ assume i x ⟨⟨x', hx'⟩, _, heq⟩, heq ▸ hx') - this, - have f ⁻¹' s = (⋃i, (coe : {x | c i x} → α) '' (f ∘ coe ⁻¹' s)), - begin - apply set.ext, - have : ∀ (x : α), f x ∈ s ↔ ∃ (i : ι), c i x ∧ f x ∈ s := - λ x, ⟨λ hx, let ⟨i, hi⟩ := h_cover x in ⟨i, hi, hx⟩, - λ ⟨i, hi, hx⟩, hx⟩, - simpa [and.comm, @and.left_comm (c _ _), ← exists_and_distrib_right], - end, - by rwa [this] - lemma closure_subtype {x : {a // p a}} {s : set {a // p a}}: x ∈ closure s ↔ (x : α) ∈ closure ((coe : _ → α) '' s) := closure_induced diff --git a/src/topology/continuous_function/basic.lean b/src/topology/continuous_function/basic.lean index 897344cee1864..fffbb5263bd4c 100644 --- a/src/topology/continuous_function/basic.lean +++ b/src/topology/continuous_function/basic.lean @@ -258,11 +258,9 @@ begin rw set.mem_Union, obtain ⟨i, hi⟩ := hS x, exact ⟨i, mem_of_mem_nhds hi⟩ }, - refine ⟨set.lift_cover S (λ i, φ i) hφ H, continuous_subtype_nhds_cover hS _⟩, - intros i, - convert (φ i).continuous, - ext x, - exact set.lift_cover_coe x, + refine ⟨set.lift_cover S (λ i, φ i) hφ H, continuous_of_cover_nhds hS $ λ i, _⟩, + rw [continuous_on_iff_continuous_restrict], + simpa only [set.restrict, set.lift_cover_coe] using (φ i).continuous end variables {S φ hφ hS} diff --git a/src/topology/continuous_on.lean b/src/topology/continuous_on.lean index e9e6b09217110..bec903efb4b3f 100644 --- a/src/topology/continuous_on.lean +++ b/src/topology/continuous_on.lean @@ -215,9 +215,12 @@ theorem nhds_within_eq_nhds_within {a : α} {s t u : set α} 𝓝[t] a = 𝓝[u] a := by rw [nhds_within_restrict t h₀ h₁, nhds_within_restrict u h₀ h₁, h₂] +@[simp] theorem nhds_within_eq_nhds {a : α} {s : set α} : 𝓝[s] a = 𝓝 a ↔ s ∈ 𝓝 a := +by rw [nhds_within, inf_eq_left, le_principal_iff] + theorem is_open.nhds_within_eq {a : α} {s : set α} (h : is_open s) (ha : a ∈ s) : 𝓝[s] a = 𝓝 a := -inf_eq_left.2 $ le_principal_iff.2 $ is_open.mem_nhds h ha +nhds_within_eq_nhds.2 $ is_open.mem_nhds h ha lemma preimage_nhds_within_coinduced {π : α → β} {s : set β} {t : set α} {a : α} (h : a ∈ t) (ht : is_open t) @@ -232,6 +235,18 @@ theorem nhds_within_union (a : α) (s t : set α) : 𝓝[s ∪ t] a = 𝓝[s] a ⊔ 𝓝[t] a := by { delta nhds_within, rw [←inf_sup_left, sup_principal] } +theorem nhds_within_bUnion {ι} {I : set ι} (hI : I.finite) (s : ι → set α) (a : α) : + 𝓝[⋃ i ∈ I, s i] a = ⨆ i ∈ I, 𝓝[s i] a := +set.finite.induction_on hI (by simp) $ λ t T _ _ hT, + by simp only [hT, nhds_within_union, supr_insert, bUnion_insert] + +theorem nhds_within_sUnion {S : set (set α)} (hS : S.finite) (a : α) : + 𝓝[⋃₀ S] a = ⨆ s ∈ S, 𝓝[s] a := +by rw [sUnion_eq_bUnion, nhds_within_bUnion hS] + +theorem nhds_within_Union {ι} [finite ι] (s : ι → set α) (a : α) : 𝓝[⋃ i, s i] a = ⨆ i, 𝓝[s i] a := +by rw [← sUnion_range, nhds_within_sUnion (finite_range s), supr_range] + theorem nhds_within_inter (a : α) (s t : set α) : 𝓝[s ∩ t] a = 𝓝[s] a ⊓ 𝓝[t] a := by { delta nhds_within, rw [inf_left_comm, inf_assoc, inf_principal, ←inf_assoc, inf_idem] } @@ -244,6 +259,10 @@ theorem nhds_within_inter_of_mem {a : α} {s t : set α} (h : s ∈ 𝓝[t] a) : 𝓝[s ∩ t] a = 𝓝[t] a := by { rw [nhds_within_inter, inf_eq_right], exact nhds_within_le_of_mem h } +theorem nhds_within_inter_of_mem' {a : α} {s t : set α} (h : s ∈ 𝓝[t] a) : + 𝓝[t ∩ s] a = 𝓝[t] a := +by rw [inter_comm, nhds_within_inter_of_mem h] + @[simp] theorem nhds_within_singleton (a : α) : 𝓝[{a}] a = pure a := by rw [nhds_within, principal_singleton, inf_eq_right.2 (pure_le_nhds a)] @@ -599,6 +618,12 @@ lemma continuous_on.prod_map {f : α → γ} {g : β → δ} {s : set α} {t : s continuous_on (prod.map f g) (s ×ˢ t) := λ ⟨x, y⟩ ⟨hx, hy⟩, continuous_within_at.prod_map (hf x hx) (hg y hy) +lemma continuous_of_cover_nhds {ι : Sort*} {f : α → β} {s : ι → set α} + (hs : ∀ x : α, ∃ i, s i ∈ 𝓝 x) (hf : ∀ i, continuous_on f (s i)) : + continuous f := +continuous_iff_continuous_at.mpr $ λ x, let ⟨i, hi⟩ := hs x in + by { rw [continuous_at, ← nhds_within_eq_nhds.2 hi], exact hf _ _ (mem_of_mem_nhds hi) } + lemma continuous_on_empty (f : α → β) : continuous_on f ∅ := λ x, false.elim diff --git a/src/topology/locally_finite.lean b/src/topology/locally_finite.lean index 88b9bd2711022..5d11f12bb8dba 100644 --- a/src/topology/locally_finite.lean +++ b/src/topology/locally_finite.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ -import topology.basic +import topology.continuous_on import order.filter.small_sets /-! @@ -69,6 +69,50 @@ lemma exists_mem_basis {ι' : Sort*} (hf : locally_finite f) {p : ι' → Prop} let ⟨i, hpi, hi⟩ := hb.small_sets.eventually_iff.mp (hf.eventually_small_sets x) in ⟨i, hpi, hi subset.rfl⟩ +protected theorem nhds_within_Union (hf : locally_finite f) (a : X) : + 𝓝[⋃ i, f i] a = ⨆ i, 𝓝[f i] a := +begin + rcases hf a with ⟨U, haU, hfin⟩, + refine le_antisymm _ (supr_le $ λ i, nhds_within_mono _ (subset_Union _ _)), + calc 𝓝[⋃ i, f i] a = 𝓝[⋃ i, f i ∩ U] a : + by rw [← Union_inter, ← nhds_within_inter_of_mem' (nhds_within_le_nhds haU)] + ... = 𝓝[⋃ i ∈ {j | (f j ∩ U).nonempty}, (f i ∩ U)] a : + by simp only [mem_set_of_eq, Union_nonempty_self] + ... = ⨆ i ∈ {j | (f j ∩ U).nonempty}, 𝓝[f i ∩ U] a : + nhds_within_bUnion hfin _ _ + ... ≤ ⨆ i, 𝓝[f i ∩ U] a : supr₂_le_supr _ _ + ... ≤ ⨆ i, 𝓝[f i] a : supr_mono (λ i, nhds_within_mono _ $ inter_subset_left _ _) +end + +lemma continuous_on_Union' {g : X → Y} (hf : locally_finite f) + (hc : ∀ i x, x ∈ closure (f i) → continuous_within_at g (f i) x) : + continuous_on g (⋃ i, f i) := +begin + rintro x -, + rw [continuous_within_at, hf.nhds_within_Union, tendsto_supr], + intro i, + by_cases hx : x ∈ closure (f i), + { exact hc i _ hx }, + { rw [mem_closure_iff_nhds_within_ne_bot, not_ne_bot] at hx, + rw [hx], + exact tendsto_bot } +end + +lemma continuous_on_Union {g : X → Y} (hf : locally_finite f) (h_cl : ∀ i, is_closed (f i)) + (h_cont : ∀ i, continuous_on g (f i)) : + continuous_on g (⋃ i, f i) := +hf.continuous_on_Union' $ λ i x hx, h_cont i x $ (h_cl i).closure_subset hx + +protected lemma continuous' {g : X → Y} (hf : locally_finite f) (h_cov : (⋃ i, f i) = univ) + (hc : ∀ i x, x ∈ closure (f i) → continuous_within_at g (f i) x) : + continuous g := +continuous_iff_continuous_on_univ.2 $ h_cov ▸ hf.continuous_on_Union' hc + +protected lemma continuous {g : X → Y} (hf : locally_finite f) (h_cov : (⋃ i, f i) = univ) + (h_cl : ∀ i, is_closed (f i)) (h_cont : ∀ i, continuous_on g (f i)) : + continuous g := +continuous_iff_continuous_on_univ.2 $ h_cov ▸ hf.continuous_on_Union h_cl h_cont + protected lemma closure (hf : locally_finite f) : locally_finite (λ i, closure (f i)) := begin intro x, @@ -78,26 +122,15 @@ begin (inter_subset_inter_right _ interior_subset) end -lemma is_closed_Union (hf : locally_finite f) (hc : ∀i, is_closed (f i)) : - is_closed (⋃i, f i) := +lemma closure_Union (h : locally_finite f) : closure (⋃ i, f i) = ⋃ i, closure (f i) := begin - simp only [← is_open_compl_iff, compl_Union, is_open_iff_mem_nhds, mem_Inter], - intros a ha, - replace ha : ∀ i, (f i)ᶜ ∈ 𝓝 a := λ i, (hc i).is_open_compl.mem_nhds (ha i), - rcases hf a with ⟨t, h_nhds, h_fin⟩, - have : t ∩ (⋂ i ∈ {i | (f i ∩ t).nonempty}, (f i)ᶜ) ∈ 𝓝 a, - from inter_mem h_nhds ((bInter_mem h_fin).2 (λ i _, ha i)), - filter_upwards [this], - simp only [mem_inter_iff, mem_Inter], - rintros b ⟨hbt, hn⟩ i hfb, - exact hn i ⟨b, hfb, hbt⟩ hfb, + ext x, + simp only [mem_closure_iff_nhds_within_ne_bot, h.nhds_within_Union, supr_ne_bot, mem_Union] end -lemma closure_Union (h : locally_finite f) : closure (⋃ i, f i) = ⋃ i, closure (f i) := -subset.antisymm - (closure_minimal (Union_mono $ λ _, subset_closure) $ - h.closure.is_closed_Union $ λ _, is_closed_closure) - (Union_subset $ λ i, closure_mono $ subset_Union _ _) +lemma is_closed_Union (hf : locally_finite f) (hc : ∀ i, is_closed (f i)) : + is_closed (⋃ i, f i) := +by simp only [← closure_eq_iff_is_closed, hf.closure_Union, (hc _).closure_eq] /-- If `f : β → set α` is a locally finite family of closed sets, then for any `x : α`, the intersection of the complements to `f i`, `x ∉ f i`, is a neighbourhood of `x`. -/ diff --git a/src/topology/subset_properties.lean b/src/topology/subset_properties.lean index d7200e25c410c..d37fdbf1062bc 100644 --- a/src/topology/subset_properties.lean +++ b/src/topology/subset_properties.lean @@ -9,6 +9,7 @@ import data.finset.order import data.set.accumulate import data.set.bool_indicator import topology.bornology.basic +import topology.locally_finite import order.minimal /-! From e8da5f215e815d9ed3455f0216ef52b53e05438a Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sun, 26 Mar 2023 01:05:13 +0000 Subject: [PATCH 0692/1291] chore(topology/basic): backport another generalization to `Sort*` (#18660) * Generalize `is_closed_Union` to `Sort*`. * Drop `is_open_Inter_prop` and `is_closed_Union_prop`. --- src/topology/basic.lean | 14 ++------------ 1 file changed, 2 insertions(+), 12 deletions(-) diff --git a/src/topology/basic.lean b/src/topology/basic.lean index da690183da981..2eb7568ea7d7c 100644 --- a/src/topology/basic.lean +++ b/src/topology/basic.lean @@ -146,10 +146,6 @@ finite.induction_on hs lemma is_open_Inter [finite ι] {s : ι → set α} (h : ∀ i, is_open (s i)) : is_open (⋂ i, s i) := is_open_sInter (finite_range _) (forall_range_iff.2 h) -lemma is_open_Inter_prop {p : Prop} {s : p → set α} - (h : ∀ h : p, is_open (s h)) : is_open (Inter s) := -by by_cases p; simp * - lemma is_open_bInter_finset {s : finset β} {f : β → set α} (h : ∀ i ∈ s, is_open (f i)) : is_open (⋂ i ∈ s, f i) := is_open_bInter (to_finite _) h @@ -212,15 +208,9 @@ finite.induction_on hs (λ a s has hs ih h, by rw bUnion_insert; exact is_closed.union (h a (mem_insert _ _)) (ih (λ i hi, h i (mem_insert_of_mem _ hi)))) -lemma is_closed_Union [finite β] {s : β → set α} (h : ∀ i, is_closed (s i)) : +lemma is_closed_Union [finite ι] {s : ι → set α} (h : ∀ i, is_closed (s i)) : is_closed (⋃ i, s i) := -suffices is_closed (⋃ (i : β) (hi : i ∈ @univ β), s i), - by convert this; simp [set.ext_iff], -is_closed_bUnion finite_univ (λ i _, h i) - -lemma is_closed_Union_prop {p : Prop} {s : p → set α} - (h : ∀ h : p, is_closed (s h)) : is_closed (Union s) := -by by_cases p; simp * +by { simp only [← is_open_compl_iff, compl_Union] at *, exact is_open_Inter h } lemma is_closed_imp {p q : α → Prop} (hp : is_open {x | p x}) (hq : is_closed {x | q x}) : is_closed {x | p x → q x} := From 76de8ae01554c3b37d66544866659ff174e66e1f Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sun, 26 Mar 2023 10:46:16 +0000 Subject: [PATCH 0693/1291] chore(*): Fix mistakes (#18654) Fix naming errors and non-defeq diamonds recently introduced. Those were discovered during the port. --- src/algebra/field/opposite.lean | 3 ++- src/algebra/group/opposite.lean | 5 ++++- src/algebra/ring/defs.lean | 4 ++-- src/algebra/ring/opposite.lean | 5 +++-- src/data/set/intervals/ord_connected.lean | 2 +- src/number_theory/multiplicity.lean | 1 + 6 files changed, 13 insertions(+), 7 deletions(-) diff --git a/src/algebra/field/opposite.lean b/src/algebra/field/opposite.lean index ef1e470581d5a..bb6e07c2a8747 100644 --- a/src/algebra/field/opposite.lean +++ b/src/algebra/field/opposite.lean @@ -53,7 +53,8 @@ instance [division_semiring α] : division_semiring αᵃᵒᵖ := { ..add_opposite.group_with_zero α, ..add_opposite.semiring α } instance [division_ring α] : division_ring αᵃᵒᵖ := -{ ..add_opposite.group_with_zero α, ..add_opposite.ring α } +{ rat_cast_mk := λ a b hb h, by rw ←div_eq_mul_inv; exact congr_arg op (rat.cast_def _), + ..add_opposite.ring α, ..add_opposite.group_with_zero α, ..add_opposite.has_rat_cast α } instance [semifield α] : semifield αᵃᵒᵖ := { ..add_opposite.division_semiring α, ..add_opposite.comm_semiring α } diff --git a/src/algebra/group/opposite.lean b/src/algebra/group/opposite.lean index f5c4d12e8848f..18f3d07d6c75f 100644 --- a/src/algebra/group/opposite.lean +++ b/src/algebra/group/opposite.lean @@ -257,7 +257,10 @@ instance [add_comm_monoid_with_one α] : add_comm_monoid_with_one αᵃᵒᵖ := ..add_opposite.add_comm_monoid α, ..add_opposite.has_one, ..add_opposite.has_nat_cast _ } instance [add_comm_group_with_one α] : add_comm_group_with_one αᵃᵒᵖ := -{ ..add_opposite.add_comm_monoid_with_one _, ..add_opposite.add_comm_group α } +{ int_cast_of_nat := λ n, congr_arg op $ int.cast_of_nat n, + int_cast_neg_succ_of_nat := λ _, congr_arg op $ int.cast_neg_succ_of_nat _, + ..add_opposite.add_comm_monoid_with_one _, ..add_opposite.add_comm_group α, + ..add_opposite.has_int_cast α } variable {α} diff --git a/src/algebra/ring/defs.lean b/src/algebra/ring/defs.lean index 08afb5f438d3d..fe52249fb3230 100644 --- a/src/algebra/ring/defs.lean +++ b/src/algebra/ring/defs.lean @@ -292,9 +292,9 @@ class non_unital_ring (α : Type*) extends non_unital_non_assoc_ring α, non_unital_semiring α /-- A unital but not-necessarily-associative ring. -/ -@[protect_proj, ancestor non_unital_non_assoc_ring non_assoc_semiring] +@[protect_proj, ancestor non_unital_non_assoc_ring non_assoc_semiring add_comm_group_with_one] class non_assoc_ring (α : Type*) extends - non_unital_non_assoc_ring α, non_assoc_semiring α, add_group_with_one α + non_unital_non_assoc_ring α, non_assoc_semiring α, add_comm_group_with_one α /-- A ring is a type with the following structures: additive commutative group (`add_comm_group`), multiplicative monoid (`monoid`), and distributive laws (`distrib`). Equivalently, a ring is a diff --git a/src/algebra/ring/opposite.lean b/src/algebra/ring/opposite.lean index 601e9da6b3baf..79f69ea9d9d64 100644 --- a/src/algebra/ring/opposite.lean +++ b/src/algebra/ring/opposite.lean @@ -144,10 +144,11 @@ instance [non_unital_ring α] : non_unital_ring αᵃᵒᵖ := .. add_opposite.distrib α} instance [non_assoc_ring α] : non_assoc_ring αᵃᵒᵖ := -{ .. add_opposite.add_comm_group α, .. add_opposite.mul_zero_one_class α, .. add_opposite.distrib α} +{ .. add_opposite.add_comm_group_with_one α, .. add_opposite.mul_zero_one_class α, + .. add_opposite.distrib α} instance [ring α] : ring αᵃᵒᵖ := -{ .. add_opposite.add_comm_group α, .. add_opposite.monoid α, .. add_opposite.semiring α } +{ .. add_opposite.non_assoc_ring α, .. add_opposite.semiring α } instance [non_unital_comm_ring α] : non_unital_comm_ring αᵃᵒᵖ := { .. add_opposite.non_unital_ring α, .. add_opposite.non_unital_comm_semiring α } diff --git a/src/data/set/intervals/ord_connected.lean b/src/data/set/intervals/ord_connected.lean index 2a7537495a4e3..b0f3f1bec8817 100644 --- a/src/data/set/intervals/ord_connected.lean +++ b/src/data/set/intervals/ord_connected.lean @@ -177,7 +177,7 @@ end preorder section partial_order variables {α : Type*} [partial_order α] {s : set α} -protected lemma is_antichain.ord_connected (hs : is_antichain (≤) s) : s.ord_connected := +protected lemma _root_.is_antichain.ord_connected (hs : is_antichain (≤) s) : s.ord_connected := ⟨λ x hx y hy z hz, by { obtain rfl := hs.eq hx hy (hz.1.trans hz.2), rw [Icc_self, mem_singleton_iff] at hz, rwa hz }⟩ diff --git a/src/number_theory/multiplicity.lean b/src/number_theory/multiplicity.lean index 82c6eed3af5f6..8df0b65f91b6c 100644 --- a/src/number_theory/multiplicity.lean +++ b/src/number_theory/multiplicity.lean @@ -252,6 +252,7 @@ begin push_cast, rw [← map_int_cast (zmod.cast_hom (show 2 ∣ 4, by norm_num) (zmod 2)) x] at hx, set y : zmod 4 := x, + change zmod.cast_hom _ (zmod 2) y = _ at hx, -- Now we can just consider each of the 4 possible values for y fin_cases y using hy; rw hy at ⊢ hx; revert hx; dec_trivial From c227d107bbada5d0d9d20287e3282c0a7f1651a0 Mon Sep 17 00:00:00 2001 From: negiizhao Date: Sun, 26 Mar 2023 13:57:30 +0000 Subject: [PATCH 0694/1291] chore(data/set/pairwise): split (#17880) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR will split most of the lemmas in `data.set.pairwise` which are independent of the `data.set.lattice`. It makes a lot of files no longer depend on `data.set.lattice`. [Zulip](https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/port.20progress/near/315072418) mathlib4 PR: https://github.com/leanprover-community/mathlib4/pull/1184 Co-authored-by: Yaël Dillies --- src/algebra/big_operators/basic.lean | 3 +- src/data/finset/option.lean | 1 - src/data/list/nodup.lean | 2 +- src/data/multiset/finset_ops.lean | 4 + src/data/prod/tprod.lean | 2 +- src/data/set/intervals/group.lean | 2 +- .../{pairwise.lean => pairwise/basic.lean} | 95 +------------- src/data/set/pairwise/lattice.lean | 123 ++++++++++++++++++ src/data/set/pointwise/smul.lean | 2 +- src/order/antichain.lean | 5 +- src/order/chain.lean | 3 +- src/order/succ_pred/interval_succ.lean | 2 +- 12 files changed, 141 insertions(+), 103 deletions(-) rename src/data/set/{pairwise.lean => pairwise/basic.lean} (79%) create mode 100644 src/data/set/pairwise/lattice.lean diff --git a/src/algebra/big_operators/basic.lean b/src/algebra/big_operators/basic.lean index 6ff8391e6c7db..c84cebf6727e3 100644 --- a/src/algebra/big_operators/basic.lean +++ b/src/algebra/big_operators/basic.lean @@ -3,7 +3,6 @@ Copyright (c) 2017 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl -/ - import algebra.big_operators.multiset.lemmas import algebra.group.pi import algebra.group_power.lemmas @@ -13,7 +12,7 @@ import data.finset.sum import data.fintype.basic import data.finset.sigma import data.multiset.powerset -import data.set.pairwise +import data.set.pairwise.basic /-! # Big operators diff --git a/src/data/finset/option.lean b/src/data/finset/option.lean index d63ce497a8d5a..88793ad39b3f3 100644 --- a/src/data/finset/option.lean +++ b/src/data/finset/option.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov, Mario Carneiro, Sean Leather -/ import data.finset.card -import order.hom.basic /-! # Finite sets in `option α` diff --git a/src/data/list/nodup.lean b/src/data/list/nodup.lean index 1b0a1b4365216..3b323391f18c2 100644 --- a/src/data/list/nodup.lean +++ b/src/data/list/nodup.lean @@ -6,7 +6,7 @@ Authors: Mario Carneiro, Kenny Lau import data.list.lattice import data.list.pairwise import data.list.forall2 -import data.set.pairwise +import data.set.pairwise.basic /-! # Lists with no duplicates diff --git a/src/data/multiset/finset_ops.lean b/src/data/multiset/finset_ops.lean index 4cc64ca2a3365..1935e66db4a65 100644 --- a/src/data/multiset/finset_ops.lean +++ b/src/data/multiset/finset_ops.lean @@ -210,3 +210,7 @@ theorem ndinter_eq_zero_iff_disjoint {s t : multiset α} : ndinter s t = 0 ↔ d by rw ← subset_zero; simp [subset_iff, disjoint] end multiset + +-- Assert that we define `finset` without the material on the set lattice. +-- Note that we cannot put this in `data.finset.basic` because we proved relevant lemmas there. +assert_not_exists set.sInter diff --git a/src/data/prod/tprod.lean b/src/data/prod/tprod.lean index e7bcd1d83256b..65718a6766ea0 100644 --- a/src/data/prod/tprod.lean +++ b/src/data/prod/tprod.lean @@ -144,7 +144,7 @@ end lemma elim_preimage_pi [decidable_eq ι] {l : list ι} (hnd : l.nodup) (h : ∀ i, i ∈ l) (t : Π i, set (α i)) : tprod.elim' h ⁻¹' pi univ t = set.tprod l t := begin - have : { i | i ∈ l} = univ, { ext i, simp [h] }, + have : { i | i ∈ l } = univ, { ext i, simp [h] }, rw [← this, ← mk_preimage_tprod, preimage_preimage], convert preimage_id, simp [tprod.mk_elim hnd h, id_def] end diff --git a/src/data/set/intervals/group.lean b/src/data/set/intervals/group.lean index b72182b7ff23f..218d2dcc821fe 100644 --- a/src/data/set/intervals/group.lean +++ b/src/data/set/intervals/group.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro, Patrick Massot, Yury Kudryashov, Rémy Degenne -/ import data.set.intervals.basic -import data.set.pairwise +import data.set.pairwise.basic import algebra.order.group.abs import algebra.group_power.lemmas diff --git a/src/data/set/pairwise.lean b/src/data/set/pairwise/basic.lean similarity index 79% rename from src/data/set/pairwise.lean rename to src/data/set/pairwise/basic.lean index 510c54cb4e6cc..04aa55cfed3c7 100644 --- a/src/data/set/pairwise.lean +++ b/src/data/set/pairwise/basic.lean @@ -3,9 +3,9 @@ Copyright (c) 2017 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl -/ +import data.set.function import logic.relation import logic.pairwise -import data.set.lattice /-! # Relations holding pairwise @@ -16,7 +16,7 @@ import data.set.lattice This file develops pairwise relations and defines pairwise disjoint indexed sets. We also prove many basic facts about `pairwise`. It is possible that an intermediate file, -with more imports than `logic.pairwise` but not importing `data.set.lattice` would be appropriate +with more imports than `logic.pairwise` but not importing `data.set.function` would be appropriate to hold many of these basic facts. ## Main declarations @@ -56,8 +56,6 @@ lemma pairwise_disjoint.mono [semilattice_inf α] [order_bot α] (hs : pairwise (disjoint on f)) (h : g ≤ f) : pairwise (disjoint on g) := hs.mono (λ i j hij, disjoint.mono (h i) (h j) hij) -alias function.injective_iff_pairwise_ne ↔ function.injective.pairwise_ne _ - namespace set lemma pairwise.mono (h : t ⊆ s) (hs : s.pairwise r) : t.pairwise r := @@ -188,23 +186,6 @@ lemma inj_on.pairwise_image {s : set ι} (h : s.inj_on f) : (f '' s).pairwise r ↔ s.pairwise (r on f) := by simp [h.eq_iff, set.pairwise] {contextual := tt} -lemma pairwise_Union {f : ι → set α} (h : directed (⊆) f) : - (⋃ n, f n).pairwise r ↔ ∀ n, (f n).pairwise r := -begin - split, - { assume H n, - exact pairwise.mono (subset_Union _ _) H }, - { assume H i hi j hj hij, - rcases mem_Union.1 hi with ⟨m, hm⟩, - rcases mem_Union.1 hj with ⟨n, hn⟩, - rcases h m n with ⟨p, mp, np⟩, - exact H p (mp hm) (np hn) hij } -end - -lemma pairwise_sUnion {r : α → α → Prop} {s : set (set α)} (h : directed_on (⊆) s) : - (⋃₀ s).pairwise r ↔ (∀ a ∈ s, set.pairwise a r) := -by { rw [sUnion_eq_Union, pairwise_Union (h.directed_coe), set_coe.forall], refl } - end set end pairwise @@ -290,14 +271,6 @@ lemma pairwise_disjoint.union (hs : s.pairwise_disjoint f) (ht : t.pairwise_disj (s ∪ t).pairwise_disjoint f := pairwise_disjoint_union.2 ⟨hs, ht, h⟩ -lemma pairwise_disjoint_Union {g : ι' → set ι} (h : directed (⊆) g) : - (⋃ n, g n).pairwise_disjoint f ↔ ∀ ⦃n⦄, (g n).pairwise_disjoint f := -pairwise_Union h - -lemma pairwise_disjoint_sUnion {s : set (set ι)} (h : directed_on (⊆) s) : - (⋃₀ s).pairwise_disjoint f ↔ ∀ ⦃a⦄, a ∈ s → set.pairwise_disjoint a f := -pairwise_sUnion h - -- classical lemma pairwise_disjoint.elim (hs : s.pairwise_disjoint f) {i j : ι} (hi : i ∈ s) (hj : j ∈ s) (h : ¬ disjoint (f i) (f j)) : @@ -321,27 +294,6 @@ hs.elim' hi hj $ λ h, hf $ (inf_of_le_left hij).symm.trans h end semilattice_inf_bot -section complete_lattice -variables [complete_lattice α] - -/-- Bind operation for `set.pairwise_disjoint`. If you want to only consider finsets of indices, you -can use `set.pairwise_disjoint.bUnion_finset`. -/ -lemma pairwise_disjoint.bUnion {s : set ι'} {g : ι' → set ι} {f : ι → α} - (hs : s.pairwise_disjoint (λ i' : ι', ⨆ i ∈ g i', f i)) - (hg : ∀ i ∈ s, (g i).pairwise_disjoint f) : - (⋃ i ∈ s, g i).pairwise_disjoint f := -begin - rintro a ha b hb hab, - simp_rw set.mem_Union at ha hb, - obtain ⟨c, hc, ha⟩ := ha, - obtain ⟨d, hd, hb⟩ := hb, - obtain hcd | hcd := eq_or_ne (g c) (g d), - { exact hg d hd (hcd.subst ha) hb hab }, - { exact (hs hc hd $ ne_of_apply_ne _ hcd).mono (le_supr₂ a ha) (le_supr₂ b hb) } -end - -end complete_lattice - /-! ### Pairwise disjoint set of sets -/ lemma pairwise_disjoint_range_singleton : @@ -359,22 +311,6 @@ lemma pairwise_disjoint.elim_set {s : set ι} {f : ι → set α} (hs : s.pairwi (hi : i ∈ s) (hj : j ∈ s) (a : α) (hai : a ∈ f i) (haj : a ∈ f j) : i = j := hs.elim hi hj $ not_disjoint_iff.2 ⟨a, hai, haj⟩ -lemma bUnion_diff_bUnion_eq {s t : set ι} {f : ι → set α} (h : (s ∪ t).pairwise_disjoint f) : - (⋃ i ∈ s, f i) \ (⋃ i ∈ t, f i) = (⋃ i ∈ s \ t, f i) := -begin - refine (bUnion_diff_bUnion_subset f s t).antisymm - (Union₂_subset $ λ i hi a ha, (mem_diff _).2 ⟨mem_bUnion hi.1 ha, _⟩), - rw mem_Union₂, rintro ⟨j, hj, haj⟩, - exact (h (or.inl hi.1) (or.inr hj) (ne_of_mem_of_not_mem hj hi.2).symm).le_bot ⟨ha, haj⟩, -end - -/-- Equivalence between a disjoint bounded union and a dependent sum. -/ -noncomputable def bUnion_eq_sigma_of_disjoint {s : set ι} {f : ι → set α} - (h : s.pairwise_disjoint f) : - (⋃ i ∈ s, f i) ≃ (Σ i : s, f i) := -(equiv.set_congr (bUnion_eq_Union _ _)).trans $ Union_eq_sigma_of_disjoint $ - λ ⟨i, hi⟩ ⟨j, hj⟩ ne, h hi hj $ λ eq, ne $ subtype.eq eq - /-- The partial images of a binary function `f` whose partial evaluations are injective are pairwise disjoint iff `f` is injective . -/ lemma pairwise_disjoint_image_right_iff {f : α → β → γ} {s : set α} {t : set β} @@ -413,30 +349,3 @@ end set lemma pairwise_disjoint_fiber (f : ι → α) : pairwise (disjoint on (λ a : α, f ⁻¹' {a})) := set.pairwise_univ.1 $ set.pairwise_disjoint_fiber f univ - - -section -variables {f : ι → set α} {s t : set ι} - -lemma set.pairwise_disjoint.subset_of_bUnion_subset_bUnion (h₀ : (s ∪ t).pairwise_disjoint f) - (h₁ : ∀ i ∈ s, (f i).nonempty) (h : (⋃ i ∈ s, f i) ⊆ ⋃ i ∈ t, f i) : - s ⊆ t := -begin - rintro i hi, - obtain ⟨a, hai⟩ := h₁ i hi, - obtain ⟨j, hj, haj⟩ := mem_Union₂.1 (h $ mem_Union₂_of_mem hi hai), - rwa h₀.eq (subset_union_left _ _ hi) (subset_union_right _ _ hj) - (not_disjoint_iff.2 ⟨a, hai, haj⟩), -end - -lemma pairwise.subset_of_bUnion_subset_bUnion (h₀ : pairwise (disjoint on f)) - (h₁ : ∀ i ∈ s, (f i).nonempty) (h : (⋃ i ∈ s, f i) ⊆ ⋃ i ∈ t, f i) : - s ⊆ t := -set.pairwise_disjoint.subset_of_bUnion_subset_bUnion (h₀.set_pairwise _) h₁ h - -lemma pairwise.bUnion_injective (h₀ : pairwise (disjoint on f)) (h₁ : ∀ i, (f i).nonempty) : - injective (λ s : set ι, ⋃ i ∈ s, f i) := -λ s t h, (h₀.subset_of_bUnion_subset_bUnion (λ _ _, h₁ _) $ h.subset).antisymm $ - h₀.subset_of_bUnion_subset_bUnion (λ _ _, h₁ _) $ h.superset - -end diff --git a/src/data/set/pairwise/lattice.lean b/src/data/set/pairwise/lattice.lean new file mode 100644 index 0000000000000..46c97b39af121 --- /dev/null +++ b/src/data/set/pairwise/lattice.lean @@ -0,0 +1,123 @@ +/- +Copyright (c) 2017 Johannes Hölzl. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johannes Hölzl +-/ +import data.set.lattice +import data.set.pairwise.basic + +/-! +# Relations holding pairwise + +In this file we prove many facts about `pairwise` and the set lattice. +-/ + +open set function + +variables {α β γ ι ι' : Type*} {r p q : α → α → Prop} + +section pairwise +variables {f g : ι → α} {s t u : set α} {a b : α} + +namespace set + +lemma pairwise_Union {f : ι → set α} (h : directed (⊆) f) : + (⋃ n, f n).pairwise r ↔ ∀ n, (f n).pairwise r := +begin + split, + { assume H n, + exact pairwise.mono (subset_Union _ _) H }, + { assume H i hi j hj hij, + rcases mem_Union.1 hi with ⟨m, hm⟩, + rcases mem_Union.1 hj with ⟨n, hn⟩, + rcases h m n with ⟨p, mp, np⟩, + exact H p (mp hm) (np hn) hij } +end + +lemma pairwise_sUnion {r : α → α → Prop} {s : set (set α)} (h : directed_on (⊆) s) : + (⋃₀ s).pairwise r ↔ (∀ a ∈ s, set.pairwise a r) := +by { rw [sUnion_eq_Union, pairwise_Union (h.directed_coe), set_coe.forall], refl } + +end set + +end pairwise + +namespace set +section partial_order_bot +variables [partial_order α] [order_bot α] {s t : set ι} {f g : ι → α} + +lemma pairwise_disjoint_Union {g : ι' → set ι} (h : directed (⊆) g) : + (⋃ n, g n).pairwise_disjoint f ↔ ∀ ⦃n⦄, (g n).pairwise_disjoint f := +pairwise_Union h + +lemma pairwise_disjoint_sUnion {s : set (set ι)} (h : directed_on (⊆) s) : + (⋃₀ s).pairwise_disjoint f ↔ ∀ ⦃a⦄, a ∈ s → set.pairwise_disjoint a f := +pairwise_sUnion h + +end partial_order_bot + +section complete_lattice +variables [complete_lattice α] + +/-- Bind operation for `set.pairwise_disjoint`. If you want to only consider finsets of indices, you +can use `set.pairwise_disjoint.bUnion_finset`. -/ +lemma pairwise_disjoint.bUnion {s : set ι'} {g : ι' → set ι} {f : ι → α} + (hs : s.pairwise_disjoint (λ i' : ι', ⨆ i ∈ g i', f i)) + (hg : ∀ i ∈ s, (g i).pairwise_disjoint f) : + (⋃ i ∈ s, g i).pairwise_disjoint f := +begin + rintro a ha b hb hab, + simp_rw set.mem_Union at ha hb, + obtain ⟨c, hc, ha⟩ := ha, + obtain ⟨d, hd, hb⟩ := hb, + obtain hcd | hcd := eq_or_ne (g c) (g d), + { exact hg d hd (hcd.subst ha) hb hab }, + { exact (hs hc hd $ ne_of_apply_ne _ hcd).mono (le_supr₂ a ha) (le_supr₂ b hb) } +end + +end complete_lattice + +lemma bUnion_diff_bUnion_eq {s t : set ι} {f : ι → set α} (h : (s ∪ t).pairwise_disjoint f) : + (⋃ i ∈ s, f i) \ (⋃ i ∈ t, f i) = (⋃ i ∈ s \ t, f i) := +begin + refine (bUnion_diff_bUnion_subset f s t).antisymm + (Union₂_subset $ λ i hi a ha, (mem_diff _).2 ⟨mem_bUnion hi.1 ha, _⟩), + rw mem_Union₂, rintro ⟨j, hj, haj⟩, + exact (h (or.inl hi.1) (or.inr hj) (ne_of_mem_of_not_mem hj hi.2).symm).le_bot ⟨ha, haj⟩, +end + +/-- Equivalence between a disjoint bounded union and a dependent sum. -/ +noncomputable def bUnion_eq_sigma_of_disjoint {s : set ι} {f : ι → set α} + (h : s.pairwise_disjoint f) : + (⋃ i ∈ s, f i) ≃ (Σ i : s, f i) := +(equiv.set_congr (bUnion_eq_Union _ _)).trans $ Union_eq_sigma_of_disjoint $ + λ ⟨i, hi⟩ ⟨j, hj⟩ ne, h hi hj $ λ eq, ne $ subtype.eq eq + +end set + + +section +variables {f : ι → set α} {s t : set ι} + +lemma set.pairwise_disjoint.subset_of_bUnion_subset_bUnion (h₀ : (s ∪ t).pairwise_disjoint f) + (h₁ : ∀ i ∈ s, (f i).nonempty) (h : (⋃ i ∈ s, f i) ⊆ ⋃ i ∈ t, f i) : + s ⊆ t := +begin + rintro i hi, + obtain ⟨a, hai⟩ := h₁ i hi, + obtain ⟨j, hj, haj⟩ := mem_Union₂.1 (h $ mem_Union₂_of_mem hi hai), + rwa h₀.eq (subset_union_left _ _ hi) (subset_union_right _ _ hj) + (not_disjoint_iff.2 ⟨a, hai, haj⟩), +end + +lemma pairwise.subset_of_bUnion_subset_bUnion (h₀ : pairwise (disjoint on f)) + (h₁ : ∀ i ∈ s, (f i).nonempty) (h : (⋃ i ∈ s, f i) ⊆ ⋃ i ∈ t, f i) : + s ⊆ t := +set.pairwise_disjoint.subset_of_bUnion_subset_bUnion (h₀.set_pairwise _) h₁ h + +lemma pairwise.bUnion_injective (h₀ : pairwise (disjoint on f)) (h₁ : ∀ i, (f i).nonempty) : + injective (λ s : set ι, ⋃ i ∈ s, f i) := +λ s t h, (h₀.subset_of_bUnion_subset_bUnion (λ _ _, h₁ _) $ h.subset).antisymm $ + h₀.subset_of_bUnion_subset_bUnion (λ _ _, h₁ _) $ h.superset + +end diff --git a/src/data/set/pointwise/smul.lean b/src/data/set/pointwise/smul.lean index 0357b1676fd5b..6900611e8e936 100644 --- a/src/data/set/pointwise/smul.lean +++ b/src/data/set/pointwise/smul.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin, Floris van Doorn -/ import algebra.module.basic -import data.set.pairwise +import data.set.pairwise.lattice import data.set.pointwise.basic import tactic.by_contra diff --git a/src/order/antichain.lean b/src/order/antichain.lean index a49fe4cf4beeb..0e62ef2616deb 100644 --- a/src/order/antichain.lean +++ b/src/order/antichain.lean @@ -3,7 +3,10 @@ Copyright (c) 2021 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ -import data.set.pairwise +import data.set.pairwise.basic +import order.bounds.basic +import order.directed +import order.hom.set /-! # Antichains diff --git a/src/order/chain.lean b/src/order/chain.lean index 9c0087af10389..34b01b9f53c5e 100644 --- a/src/order/chain.lean +++ b/src/order/chain.lean @@ -3,7 +3,8 @@ Copyright (c) 2017 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl -/ -import data.set.pairwise +import data.set.pairwise.basic +import data.set.lattice import data.set_like.basic /-! diff --git a/src/order/succ_pred/interval_succ.lean b/src/order/succ_pred/interval_succ.lean index 79a635165d2a4..b2a00b4e2f070 100644 --- a/src/order/succ_pred/interval_succ.lean +++ b/src/order/succ_pred/interval_succ.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ -import data.set.pairwise +import data.set.pairwise.basic import order.succ_pred.basic /-! From b67044ba53af18680e1dd246861d9584e968495d Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Sun, 26 Mar 2023 17:00:46 +0000 Subject: [PATCH 0695/1291] feat(set_theory/ordinal/arithmetic): miscellaneous arithmetic lemmas (#15990) Will be used to prove statements about the Cantor Normal Form. --- src/set_theory/ordinal/arithmetic.lean | 22 +++++++++++++++++++++- src/set_theory/ordinal/exponential.lean | 16 +++++++++++----- src/set_theory/ordinal/notation.lean | 2 +- 3 files changed, 33 insertions(+), 7 deletions(-) diff --git a/src/set_theory/ordinal/arithmetic.lean b/src/set_theory/ordinal/arithmetic.lean index fe375d7ea1343..67d8803ccf422 100644 --- a/src/set_theory/ordinal/arithmetic.lean +++ b/src/set_theory/ordinal/arithmetic.lean @@ -479,7 +479,7 @@ protected theorem sub_eq_zero_iff_le {a b : ordinal} : a - b = 0 ↔ a ≤ b := theorem sub_sub (a b c : ordinal) : a - b - c = a - (b + c) := eq_of_forall_ge_iff $ λ d, by rw [sub_le, sub_le, sub_le, add_assoc] -theorem add_sub_add_cancel (a b c : ordinal) : a + b - (a + c) = b - c := +@[simp] theorem add_sub_add_cancel (a b c : ordinal) : a + b - (a + c) = b - c := by rw [← sub_sub, add_sub_cancel] theorem sub_is_limit {a b} (l : is_limit a) (h : b < a) : is_limit (a - b) := @@ -804,6 +804,8 @@ instance : has_mod ordinal := ⟨λ a b, a - b * (a / b)⟩ theorem mod_def (a b : ordinal) : a % b = a - b * (a / b) := rfl +theorem mod_le (a b : ordinal) : a % b ≤ a := sub_le_self a _ + @[simp] theorem mod_zero (a : ordinal) : a % 0 = a := by simp only [mod_def, div_zero, zero_mul, sub_zero] @@ -841,6 +843,24 @@ end theorem dvd_iff_mod_eq_zero {a b : ordinal} : b ∣ a ↔ a % b = 0 := ⟨mod_eq_zero_of_dvd, dvd_of_mod_eq_zero⟩ +@[simp] theorem mul_add_mod_self (x y z : ordinal) : (x * y + z) % x = z % x := +begin + rcases eq_or_ne x 0 with rfl | hx, + { simp }, + { rwa [mod_def, mul_add_div, mul_add, ←sub_sub, add_sub_cancel, mod_def] } +end + +@[simp] theorem mul_mod (x y : ordinal) : x * y % x = 0 := by simpa using mul_add_mod_self x y 0 + +theorem mod_mod_of_dvd (a : ordinal) {b c : ordinal} (h : c ∣ b) : a % b % c = a % c := +begin + nth_rewrite_rhs 0 ←div_add_mod a b, + rcases h with ⟨d, rfl⟩, + rw [mul_assoc, mul_add_mod_self] +end + +@[simp] theorem mod_mod (a b : ordinal) : a % b % b = a % b := mod_mod_of_dvd a dvd_rfl + /-! ### Families of ordinals There are two kinds of indexed families that naturally arise when dealing with ordinals: those diff --git a/src/set_theory/ordinal/exponential.lean b/src/set_theory/ordinal/exponential.lean index 92c0c1bcabc28..0b1e2723c625a 100644 --- a/src/set_theory/ordinal/exponential.lean +++ b/src/set_theory/ordinal/exponential.lean @@ -190,12 +190,10 @@ end theorem opow_one_add (a b : ordinal) : a ^ (1 + b) = a * a ^ b := by rw [opow_add, opow_one] -theorem opow_dvd_opow (a) {b c : ordinal} - (h : b ≤ c) : a ^ b ∣ a ^ c := -by { rw [← ordinal.add_sub_cancel_of_le h, opow_add], apply dvd_mul_right } +theorem opow_dvd_opow (a) {b c : ordinal} (h : b ≤ c) : a ^ b ∣ a ^ c := +⟨a ^ (c - b), by rw [←opow_add, ordinal.add_sub_cancel_of_le h] ⟩ -theorem opow_dvd_opow_iff {a b c : ordinal} - (a1 : 1 < a) : a ^ b ∣ a ^ c ↔ b ≤ c := +theorem opow_dvd_opow_iff {a b c : ordinal} (a1 : 1 < a) : a ^ b ∣ a ^ c ↔ b ≤ c := ⟨λ h, le_of_not_lt $ λ hn, not_le_of_lt ((opow_lt_opow_iff_right a1).2 hn) $ le_of_dvd (opow_ne_zero _ $ one_le_iff_ne_zero.1 $ a1.le) h, @@ -385,6 +383,14 @@ begin rw [add_zero, mul_one] end +theorem div_opow_log_pos (b : ordinal) {o : ordinal} (ho : o ≠ 0) : 0 < o / b ^ log b o := +begin + rcases eq_zero_or_pos b with (rfl | hb), + { simpa using ordinal.pos_iff_ne_zero.2 ho }, + { rw div_pos (opow_ne_zero _ hb.ne'), + exact opow_log_le_self b ho } +end + theorem div_opow_log_lt {b : ordinal} (o : ordinal) (hb : 1 < b) : o / b ^ log b o < b := begin rw [div_lt (opow_pos _ (zero_lt_one.trans hb)).ne', ←opow_succ], diff --git a/src/set_theory/ordinal/notation.lean b/src/set_theory/ordinal/notation.lean index a02aae9dce46c..cb4d422862155 100644 --- a/src/set_theory/ordinal/notation.lean +++ b/src/set_theory/ordinal/notation.lean @@ -428,7 +428,7 @@ instance sub_NF (o₁ o₂) : ∀ [NF o₁] [NF o₂], NF (o₁ - o₂) { change e₁ = e₂ at ee, substI e₂, unfold sub._match_1, cases mn : (n₁:ℕ) - n₂; dsimp only [sub._match_2], { by_cases en : n₁ = n₂, - { simp [en], rwa [add_sub_add_cancel] }, + { simpa [en] }, { simp [en, -repr], exact (ordinal.sub_eq_zero_iff_le.2 $ le_of_lt $ oadd_lt_oadd_2 h₁ $ lt_of_le_of_ne (tsub_eq_zero_iff_le.1 mn) (mt pnat.eq en)).symm } }, From cea83e192eae2d368ab2b500a0975667da42c920 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sun, 26 Mar 2023 17:00:47 +0000 Subject: [PATCH 0696/1291] feat(data/finmap): add an equivalence with pairs `(keys, lookup)` (#18151) A non-dependent version of the equivalence requested [on Zulip](https://leanprover.zulipchat.com/#narrow/stream/113488-general/topic/Proving.20equality.20of.20sigma.20types). --- src/data/finmap.lean | 73 ++++++++++++++++++++++++++++++++++++++++++-- 1 file changed, 70 insertions(+), 3 deletions(-) diff --git a/src/data/finmap.lean b/src/data/finmap.lean index 0f49f8d4311e9..0518cbb1a88c7 100644 --- a/src/data/finmap.lean +++ b/src/data/finmap.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sean Leather, Mario Carneiro -/ import data.list.alist -import data.finset.basic +import data.finset.sigma import data.part /-! # Finite maps over `multiset` @@ -35,6 +35,14 @@ quot.lift_on s list.nodupkeys (λ s t p, propext $ perm_nodupkeys p) @[simp] theorem coe_nodupkeys {l : list (sigma β)} : @nodupkeys α β l ↔ l.nodupkeys := iff.rfl +lemma nodup_keys {m : multiset (Σ a, β a)} : m.keys.nodup ↔ m.nodupkeys := +by { rcases m with ⟨l⟩, refl } + +alias nodup_keys ↔ _ nodupkeys.nodup_keys + +lemma nodupkeys.nodup {m : multiset (Σ a, β a)} (h : m.nodupkeys) : m.nodup := +h.nodup_keys.of_map _ + end multiset /-! ### finmap -/ @@ -63,6 +71,8 @@ def list.to_finmap [decidable_eq α] (s : list (sigma β)) : finmap β := s.to_a namespace finmap open alist +lemma nodup_entries (f : finmap β) : f.entries.nodup := f.nodupkeys.nodup + /-! ### lifting from alist -/ /-- Lift a permutation-respecting function on `alist` to `finmap`. -/ @@ -129,7 +139,7 @@ theorem mem_def {a : α} {s : finmap β} : /-- The set of keys of a finite map. -/ def keys (s : finmap β) : finset α := -⟨s.entries.keys, induction_on s keys_nodup⟩ +⟨s.entries.keys, s.nodupkeys.nodup_keys⟩ @[simp] theorem keys_val (s : alist β) : (keys ⟦s⟧).val = s.keys := rfl @@ -197,6 +207,23 @@ induction_on s $ λ s, alist.lookup_is_some theorem lookup_eq_none {a} {s : finmap β} : lookup a s = none ↔ a ∉ s := induction_on s $ λ s, alist.lookup_eq_none +lemma mem_lookup_iff {f : finmap β} {a : α} {b : β a} : + b ∈ f.lookup a ↔ sigma.mk a b ∈ f.entries := +by { rcases f with ⟨⟨l⟩, hl⟩, exact list.mem_lookup_iff hl } + +/-- A version of `finmap.mem_lookup_iff` with LHS in the simp-normal form. -/ +lemma lookup_eq_some_iff {f : finmap β} {a : α} {b : β a} : + f.lookup a = some b ↔ sigma.mk a b ∈ f.entries := +mem_lookup_iff + +@[simp] lemma sigma_keys_lookup (f : finmap β) : + f.keys.sigma (λ i, (f.lookup i).to_finset) = ⟨f.entries, f.nodup_entries⟩ := +begin + ext x, + have : x ∈ f.entries → x.fst ∈ f.keys, from multiset.mem_map_of_mem _, + simpa [lookup_eq_some_iff] +end + @[simp] lemma lookup_singleton_eq {a : α} {b : β a} : (singleton a b).lookup a = some b := by rw [singleton, lookup_to_finmap, alist.singleton, alist.lookup, lookup_cons_eq] @@ -206,7 +233,7 @@ decidable_of_iff _ lookup_is_some lemma mem_iff {a : α} {s : finmap β} : a ∈ s ↔ ∃ b, s.lookup a = some b := induction_on s $ λ s, iff.trans list.mem_keys $ exists_congr $ λ b, -(mem_lookup_iff s.nodupkeys).symm +(list.mem_lookup_iff s.nodupkeys).symm lemma mem_of_lookup_eq_some {a : α} {b : β a} {s : finmap β} (h : s.lookup a = some b) : a ∈ s := mem_iff.mpr ⟨_, h⟩ @@ -221,6 +248,46 @@ begin rw h, end +/-- An equivalence between `finmap β` and pairs `(keys : finset α, lookup : Π a, option (β a))` such +that `(lookup a).is_some ↔ a ∈ keys`. -/ +@[simps apply_coe_fst apply_coe_snd] +def keys_lookup_equiv : + finmap β ≃ {f : finset α × (Π a, option (β a)) // ∀ i, (f.2 i).is_some ↔ i ∈ f.1} := +{ to_fun := λ f, ⟨(f.keys, λ i, f.lookup i), λ i, lookup_is_some⟩, + inv_fun := λ f, ⟨(f.1.1.sigma $ λ i, (f.1.2 i).to_finset).val, + begin + refine multiset.nodup_keys.1 ((finset.nodup _).map_on _), + simp only [finset.mem_val, finset.mem_sigma, option.mem_to_finset, option.mem_def], + rintro ⟨i, x⟩ ⟨hi, hx⟩ ⟨j, y⟩ ⟨hj, hy⟩ (rfl : i = j), + obtain rfl : x = y, from option.some.inj (hx.symm.trans hy), + refl + end⟩, + left_inv := λ f, ext $ by simp, + right_inv := λ ⟨⟨s, f⟩, hf⟩, + begin + ext : 2; dsimp [keys], + { ext1 i, + have : i ∈ s → (∃ x, f i = some x), + from λ hi, ⟨option.get _, option.get_mem $ (hf i).2 hi⟩, + simpa [multiset.keys] }, + { ext i x : 2, + simp only [option.mem_def, lookup_eq_some_iff, finset.mem_val, finset.mem_sigma, + option.mem_to_finset, and_iff_right_iff_imp, ← hf], + exact λ h, option.is_some_iff_exists.2 ⟨_, h⟩ } + end } + +@[simp] lemma keys_lookup_equiv_symm_apply_keys : + ∀ f : {f : finset α × (Π a, option (β a)) // ∀ i, (f.2 i).is_some ↔ i ∈ f.1}, + (keys_lookup_equiv.symm f).keys = (f : finset α × Π a, option (β a)).1 := +keys_lookup_equiv.surjective.forall.2 $ λ f, + by simp only [equiv.symm_apply_apply, keys_lookup_equiv_apply_coe_fst] + +@[simp] lemma keys_lookup_equiv_symm_apply_lookup : + ∀ (f : {f : finset α × (Π a, option (β a)) // ∀ i, (f.2 i).is_some ↔ i ∈ f.1}) a, + (keys_lookup_equiv.symm f).lookup a = (f : finset α × Π a, option (β a)).2 a := +keys_lookup_equiv.surjective.forall.2 $ λ f a, + by simp only [equiv.symm_apply_apply, keys_lookup_equiv_apply_coe_snd] + /-! ### replace -/ /-- Replace a key with a given value in a finite map. From ce86f4e05e9a9b8da5e316b22c76ce76440c56a1 Mon Sep 17 00:00:00 2001 From: Moritz Doll Date: Sun, 26 Mar 2023 19:10:16 +0000 Subject: [PATCH 0697/1291] feat(analysis/locally_convex/with_seminorms): convergence along filters (#18664) --- .../locally_convex/with_seminorms.lean | 33 +++++++++++++++++++ 1 file changed, 33 insertions(+) diff --git a/src/analysis/locally_convex/with_seminorms.lean b/src/analysis/locally_convex/with_seminorms.lean index 01b9e0690d245..a39d6f380a131 100644 --- a/src/analysis/locally_convex/with_seminorms.lean +++ b/src/analysis/locally_convex/with_seminorms.lean @@ -371,6 +371,39 @@ end end topology +section tendsto + +variables [normed_field 𝕜] [add_comm_group E] [module 𝕜 E] [nonempty ι] [topological_space E] +variables {p : seminorm_family 𝕜 E ι} + +/-- Convergence along filters for `with_seminorms`. + +Variant with `finset.sup`. -/ +lemma with_seminorms.tendsto_nhds' (hp : with_seminorms p) (u : F → E) {f : filter F} (y₀ : E) : + filter.tendsto u f (𝓝 y₀) ↔ ∀ (s : finset ι) ε, 0 < ε → ∀ᶠ x in f, s.sup p (u x - y₀) < ε := +by simp [hp.has_basis_ball.tendsto_right_iff] + +/-- Convergence along filters for `with_seminorms`. -/ +lemma with_seminorms.tendsto_nhds (hp : with_seminorms p) (u : F → E) {f : filter F} (y₀ : E) : + filter.tendsto u f (𝓝 y₀) ↔ ∀ i ε, 0 < ε → ∀ᶠ x in f, p i (u x - y₀) < ε := +begin + rw hp.tendsto_nhds' u y₀, + exact ⟨λ h i, by simpa only [finset.sup_singleton] using h {i}, + λ h s ε hε, (s.eventually_all.2 $ λ i _, h i ε hε).mono (λ _, finset_sup_apply_lt hε)⟩, +end + +variables [semilattice_sup F] [nonempty F] + +/-- Limit `→ ∞` for `with_seminorms`. -/ +lemma with_seminorms.tendsto_nhds_at_top (hp : with_seminorms p) (u : F → E) (y₀ : E) : + filter.tendsto u filter.at_top (𝓝 y₀) ↔ ∀ i ε, 0 < ε → ∃ x₀, ∀ x, x₀ ≤ x → p i (u x - y₀) < ε := +begin + rw hp.tendsto_nhds u y₀, + exact forall₃_congr (λ _ _ _, filter.eventually_at_top), +end + +end tendsto + section topological_add_group variables [normed_field 𝕜] [add_comm_group E] [module 𝕜 E] From c5dd93108bec0f9023919bc97bc1f68f88edb7bd Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Sun, 26 Mar 2023 23:49:07 +0000 Subject: [PATCH 0698/1291] feat(set_theory/zfc/basic): tweak Class hom lemmas (#18295) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR renames a bunch of `hom` lemmas to better match the style of the rest of mathlib, and tags them as `norm_cast`. We also add the corresponding lemmas for the union. Co-authored-by: Yaël Dillies --- src/set_theory/zfc/basic.lean | 68 +++++++++++++++++++---------------- 1 file changed, 38 insertions(+), 30 deletions(-) diff --git a/src/set_theory/zfc/basic.lean b/src/set_theory/zfc/basic.lean index 5aa704b4fa60d..e0e162b352924 100644 --- a/src/set_theory/zfc/basic.lean +++ b/src/set_theory/zfc/basic.lean @@ -998,59 +998,67 @@ Set.ext $ λ z, by { change (x : Class.{u}) z ↔ (y : Class.{u}) z, rw h } @[simp] theorem to_Set_of_Set (A : Class.{u}) (x : Set.{u}) : to_Set A x ↔ A x := ⟨λ ⟨y, yx, py⟩, by rwa of_Set.inj yx at py, λ px, ⟨x, rfl, px⟩⟩ -@[simp] theorem mem_hom_left (x : Set.{u}) (A : Class.{u}) : (x : Class.{u}) ∈ A ↔ A x := +@[simp, norm_cast] theorem coe_mem {x : Set.{u}} {A : Class.{u}} : (x : Class.{u}) ∈ A ↔ A x := to_Set_of_Set _ _ -@[simp] theorem mem_hom_right (x y : Set.{u}) : (y : Class.{u}) x ↔ x ∈ y := iff.rfl +@[simp] theorem coe_apply {x y : Set.{u}} : (y : Class.{u}) x ↔ x ∈ y := iff.rfl -@[simp] theorem subset_hom (x y : Set.{u}) : (x : Class.{u}) ⊆ y ↔ x ⊆ y := iff.rfl +@[simp, norm_cast] theorem coe_subset (x y : Set.{u}) : (x : Class.{u}) ⊆ y ↔ x ⊆ y := iff.rfl -@[simp] theorem sep_hom (p : Class.{u}) (x : Set.{u}) : +@[simp, norm_cast] theorem coe_sep (p : Class.{u}) (x : Set.{u}) : (↑{y ∈ x | p y} : Class.{u}) = {y ∈ x | p y} := set.ext $ λ y, Set.mem_sep -@[simp] theorem empty_hom : ↑(∅ : Set.{u}) = (∅ : Class.{u}) := -set.ext $ λ y, (iff_false _).2 (Set.not_mem_empty y) +@[simp, norm_cast] theorem coe_empty : ↑(∅ : Set.{u}) = (∅ : Class.{u}) := +set.ext $ λ y, (iff_false _).2 $ Set.not_mem_empty y -@[simp] theorem insert_hom (x y : Set.{u}) : (@insert Set.{u} Class.{u} _ x y) = ↑(insert x y) := -set.ext $ λ z, iff.symm Set.mem_insert_iff +@[simp, norm_cast] theorem coe_insert (x y : Set.{u}) : + ↑(insert x y) = @insert Set.{u} Class.{u} _ x y := +set.ext $ λ z, Set.mem_insert_iff -@[simp] theorem union_hom (x y : Set.{u}) : (x : Class.{u}) ∪ y = (x ∪ y : Set.{u}) := -set.ext $ λ z, iff.symm Set.mem_union +@[simp, norm_cast] theorem coe_union (x y : Set.{u}) : ↑(x ∪ y) = (x : Class.{u}) ∪ y := +set.ext $ λ z, Set.mem_union -@[simp] theorem inter_hom (x y : Set.{u}) : (x : Class.{u}) ∩ y = (x ∩ y : Set.{u}) := -set.ext $ λ z, iff.symm Set.mem_inter +@[simp, norm_cast] theorem coe_inter (x y : Set.{u}) : ↑(x ∩ y) = (x : Class.{u}) ∩ y := +set.ext $ λ z, Set.mem_inter -@[simp] theorem diff_hom (x y : Set.{u}) : (x : Class.{u}) \ y = (x \ y : Set.{u}) := -set.ext $ λ z, iff.symm Set.mem_diff +@[simp, norm_cast] theorem coe_diff (x y : Set.{u}) : ↑(x \ y) = (x : Class.{u}) \ y := +set.ext $ λ z, Set.mem_diff -@[simp] theorem powerset_hom (x : Set.{u}) : powerset.{u} x = Set.powerset x := -set.ext $ λ z, iff.symm Set.mem_powerset +@[simp, norm_cast] theorem coe_powerset (x : Set.{u}) : ↑x.powerset = powerset.{u} x := +set.ext $ λ z, Set.mem_powerset -@[simp] theorem sUnion_hom (x : Set.{u}) : ⋃₀ (x : Class.{u}) = ⋃₀ x := -set.ext $ λ z, by { refine iff.trans _ Set.mem_sUnion.symm, exact -⟨λ ⟨._, ⟨a, rfl, ax⟩, za⟩, ⟨a, ax, za⟩, λ ⟨a, ax, za⟩, ⟨_, ⟨a, rfl, ax⟩, za⟩⟩ } +@[simp] theorem powerset_apply {A : Class.{u}} {x : Set.{u}} : powerset A x ↔ ↑x ⊆ A := iff.rfl + +@[simp] theorem sUnion_apply {x : Class} {y : Set} : (⋃₀ x) y ↔ ∃ z : Set, x z ∧ y ∈ z := +begin + split, + { rintro ⟨-, ⟨z, rfl, hxz⟩, hyz⟩, + exact ⟨z, hxz, hyz⟩ }, + { exact λ ⟨z, hxz, hyz⟩, ⟨_, coe_mem.2 hxz, hyz⟩ } +end + +@[simp, norm_cast] theorem coe_sUnion (x : Set.{u}) : ↑(⋃₀ x) = ⋃₀ (x : Class.{u}) := +set.ext $ λ y, Set.mem_sUnion.trans (sUnion_apply.trans $ by simp_rw [coe_apply, exists_prop]).symm @[ext] theorem ext {x y : Class.{u}} : (∀ z : Class.{u}, z ∈ x ↔ z ∈ y) → x = y := begin refine λ h, set.ext (λ z, _), change x z ↔ y z, - rw [←mem_hom_left z x, ←mem_hom_left z y], + rw [←@coe_mem z x, ←@coe_mem z y], exact h z end theorem ext_iff {x y : Class.{u}} : x = y ↔ (∀ z : Class.{u}, z ∈ x ↔ z ∈ y) := ⟨λ h, by simp [h], ext⟩ -theorem coe_mem_powerset {x : Class.{u}} {y : Set.{u}} : powerset x y ↔ ↑y ⊆ x := iff.rfl - @[simp] theorem mem_sUnion {x y : Class.{u}} : y ∈ ⋃₀ x ↔ ∃ z, z ∈ x ∧ y ∈ z := begin split, - { rintro ⟨w, rfl, ⟨z, hzx, hwz⟩⟩, - exact ⟨z, hzx, (mem_hom_left _ _).2 hwz⟩ }, - { rintro ⟨w, hwx, ⟨z, rfl, hwz⟩⟩, - exact ⟨z, rfl, ⟨w, hwx, hwz⟩⟩ } + { rintro ⟨w, rfl, z, hzx, hwz⟩, + exact ⟨z, hzx, coe_mem.2 hwz⟩ }, + { rintro ⟨w, hwx, z, rfl, hwz⟩, + exact ⟨z, rfl, w, hwx, hwz⟩ } end @[simp] theorem sUnion_empty : ⋃₀ (∅ : Class.{u}) = (∅ : Class.{u}) := @@ -1062,7 +1070,7 @@ theorem eq_univ_of_powerset_subset {A : Class} (hA : powerset A ⊆ A) : A = uni eq_univ_of_forall begin by_contra' hnA, exact well_founded.min_mem Set.mem_wf _ hnA (hA $ λ x hx, not_not.1 $ - λ hB, well_founded.not_lt_min Set.mem_wf _ hnA hB $ (mem_hom_right _ _).1 hx) + λ hB, well_founded.not_lt_min Set.mem_wf _ hnA hB $ coe_apply.1 hx) end /-- The definite description operator, which is `{x}` if `{y | A y} = {x}` and `∅` otherwise. -/ @@ -1078,7 +1086,7 @@ set.ext $ λ y, ⟨λ ⟨._, ⟨x', rfl, h⟩, yx'⟩, by rwa ←((H x').1 $ (h theorem iota_ex (A) : iota.{u} A ∈ univ.{u} := mem_univ.2 $ or.elim (classical.em $ ∃ x, ∀ y, A y ↔ y = x) (λ ⟨x, h⟩, ⟨x, eq.symm $ iota_val A x h⟩) - (λ hn, ⟨∅, set.ext (λ z, empty_hom.symm ▸ ⟨false.rec _, λ ⟨._, ⟨x, rfl, H⟩, zA⟩, hn ⟨x, H⟩⟩)⟩) + (λ hn, ⟨∅, set.ext (λ z, coe_empty.symm ▸ ⟨false.rec _, λ ⟨._, ⟨x, rfl, H⟩, zA⟩, hn ⟨x, H⟩⟩)⟩) /-- Function value -/ def fval (F A : Class.{u}) : Class.{u} := iota (λ y, to_Set (λ x, F (Set.pair x y)) A) @@ -1093,7 +1101,7 @@ namespace Set @[simp] theorem map_fval {f : Set.{u} → Set.{u}} [H : pSet.definable 1 f] {x y : Set.{u}} (h : y ∈ x) : (Set.map f x ′ y : Class.{u}) = f y := -Class.iota_val _ _ (λ z, by { rw [Class.to_Set_of_Set, Class.mem_hom_right, mem_map], exact +Class.iota_val _ _ (λ z, by { rw [Class.to_Set_of_Set, Class.coe_apply, mem_map], exact ⟨λ ⟨w, wz, pr⟩, let ⟨wy, fw⟩ := Set.pair_injective pr in by rw[←fw, wy], λ e, by { subst e, exact ⟨_, h, rfl⟩ }⟩ }) @@ -1115,7 +1123,7 @@ theorem choice_is_func : is_func x (⋃₀ x) (choice x) := theorem choice_mem (y : Set.{u}) (yx : y ∈ x) : (choice x ′ y : Class.{u}) ∈ (y : Class.{u}) := begin delta choice, - rw [map_fval yx, Class.mem_hom_left, Class.mem_hom_right], + rw [map_fval yx, Class.coe_mem, Class.coe_apply], exact choice_mem_aux x h y yx end From 3dec44d0b621a174c56e994da4aae15ba60110a2 Mon Sep 17 00:00:00 2001 From: Amelia Livingston <101damnations@github.com> Date: Mon, 27 Mar 2023 02:38:29 +0000 Subject: [PATCH 0699/1291] feat(representation_theory/group_cohomology_resolution): add isomorphism with nth inhomogeneous cochains (#18159) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Given a $k$-linear $G$-representation $A,$ this defines the $k$-linear isomorphism between functions $G^n \to A$ and representation morphisms $Hom(k[G^{n + 1}], A),$ called `Rep.diagonal_hom_equiv`. Co-authored-by: Joël Riou --- src/category_theory/linear/basic.lean | 22 +++ src/linear_algebra/finsupp.lean | 20 ++- src/representation_theory/Action.lean | 4 +- src/representation_theory/Rep.lean | 137 +++++++++++++++++- src/representation_theory/basic.lean | 4 + .../group_cohomology_resolution.lean | 56 ++++++- 6 files changed, 236 insertions(+), 7 deletions(-) diff --git a/src/category_theory/linear/basic.lean b/src/category_theory/linear/basic.lean index 2f1106b55a277..59a42152c0844 100644 --- a/src/category_theory/linear/basic.lean +++ b/src/category_theory/linear/basic.lean @@ -130,6 +130,28 @@ instance {X Y : C} (f : X ⟶ Y) [mono f] (r : R) [invertible r] : mono (r • f simpa [smul_smul] using congr_arg (λ f, ⅟r • f) H, end⟩ +/-- Given isomorphic objects `X ≅ Y, W ≅ Z` in a `k`-linear category, we have a `k`-linear +isomorphism between `Hom(X, W)` and `Hom(Y, Z).` -/ +def hom_congr (k : Type*) {C : Type*} [category C] [semiring k] + [preadditive C] [linear k C] {X Y W Z : C} (f₁ : X ≅ Y) (f₂ : W ≅ Z) : + (X ⟶ W) ≃ₗ[k] (Y ⟶ Z) := +{ inv_fun := (left_comp k W f₁.hom).comp (right_comp k Y f₂.symm.hom), + left_inv := λ x, by simp only [iso.symm_hom, linear_map.to_fun_eq_coe, linear_map.coe_comp, + function.comp_app, left_comp_apply, right_comp_apply, category.assoc, iso.hom_inv_id, + category.comp_id, iso.hom_inv_id_assoc], + right_inv := λ x, by simp only [iso.symm_hom, linear_map.coe_comp, function.comp_app, + right_comp_apply, left_comp_apply, linear_map.to_fun_eq_coe, iso.inv_hom_id_assoc, + category.assoc, iso.inv_hom_id, category.comp_id], + ..(right_comp k Y f₂.hom).comp (left_comp k W f₁.symm.hom) } + +lemma hom_congr_apply (k : Type*) {C : Type*} [category C] [semiring k] + [preadditive C] [linear k C] {X Y W Z : C} (f₁ : X ≅ Y) (f₂ : W ≅ Z) (f : X ⟶ W) : + hom_congr k f₁ f₂ f = (f₁.inv ≫ f) ≫ f₂.hom := rfl + +lemma hom_congr_symm_apply (k : Type*) {C : Type*} [category C] [semiring k] + [preadditive C] [linear k C] {X Y W Z : C} (f₁ : X ≅ Y) (f₂ : W ≅ Z) (f : Y ⟶ Z) : + (hom_congr k f₁ f₂).symm f = f₁.hom ≫ f ≫ f₂.inv := rfl + end section diff --git a/src/linear_algebra/finsupp.lean b/src/linear_algebra/finsupp.lean index 8e38ee6586f15..0b05854234f17 100644 --- a/src/linear_algebra/finsupp.lean +++ b/src/linear_algebra/finsupp.lean @@ -344,7 +344,7 @@ theorem lsum_symm_apply (f : (α →₀ M) →ₗ[R] N) (x : α) : end lsum section -variables (M) (R) (X : Type*) +variables (M) (R) (X : Type*) (S) [module S M] [smul_comm_class R S M] /-- A slight rearrangement from `lsum` gives us @@ -362,6 +362,24 @@ lemma lift_apply (f) (g) : ((lift M R X) f) g = g.sum (λ x r, r • f x) := rfl +/-- Given compatible `S` and `R`-module structures on `M` and a type `X`, the set of functions +`X → M` is `S`-linearly equivalent to the `R`-linear maps from the free `R`-module +on `X` to `M`. -/ +noncomputable def llift : (X → M) ≃ₗ[S] ((X →₀ R) →ₗ[R] M) := +{ map_smul' := + begin + intros, + dsimp, + ext, + simp only [coe_comp, function.comp_app, lsingle_apply, lift_apply, pi.smul_apply, + sum_single_index, zero_smul, one_smul, linear_map.smul_apply], + end, ..lift M R X } + +@[simp] lemma llift_apply (f : X → M) (x : X →₀ R) : + llift M R S X f x = lift M R X f x := rfl + +@[simp] lemma llift_symm_apply (f : (X →₀ R) →ₗ[R] M) (x : X) : + (llift M R S X).symm f x = f (single x 1) := rfl end section lmap_domain diff --git a/src/representation_theory/Action.lean b/src/representation_theory/Action.lean index e20119ba8658c..bb5e99b5489d4 100644 --- a/src/representation_theory/Action.lean +++ b/src/representation_theory/Action.lean @@ -504,12 +504,12 @@ rfl rfl @[simp] lemma functor_category_monoidal_equivalence.functor_map {A B : Action V G} (f : A ⟶ B) : - (functor_category_monoidal_equivalence _ _).1.1.map f + (functor_category_monoidal_equivalence _ _).map f = functor_category_equivalence.functor.map f := rfl @[simp] lemma functor_category_monoidal_equivalence.inverse_map {A B : single_obj G ⥤ V} (f : A ⟶ B) : - (functor_category_monoidal_equivalence _ _).1.inv.map f + (functor_category_monoidal_equivalence _ _).inv.map f = functor_category_equivalence.inverse.map f := rfl variables (H : Group.{u}) diff --git a/src/representation_theory/Rep.lean b/src/representation_theory/Rep.lean index f9d6375170a47..20544f6b3366d 100644 --- a/src/representation_theory/Rep.lean +++ b/src/representation_theory/Rep.lean @@ -68,12 +68,26 @@ lemma coe_of {V : Type u} [add_comm_group V] [module k V] (ρ : G →* (V →ₗ @[simp] lemma of_ρ {V : Type u} [add_comm_group V] [module k V] (ρ : G →* (V →ₗ[k] V)) : (of ρ).ρ = ρ := rfl +@[simp] lemma Action_ρ_eq_ρ {A : Rep k G} : Action.ρ A = A.ρ := rfl + +/-- Allows us to apply lemmas about the underlying `ρ`, which would take an element `g : G` rather +than `g : Mon.of G` as an argument. -/ +lemma of_ρ_apply {V : Type u} [add_comm_group V] [module k V] + (ρ : representation k G V) (g : Mon.of G) : + (Rep.of ρ).ρ g = ρ (g : G) := rfl + -- Verify that limits are calculated correctly. noncomputable example : preserves_limits (forget₂ (Rep k G) (Module.{u} k)) := by apply_instance noncomputable example : preserves_colimits (forget₂ (Rep k G) (Module.{u} k)) := by apply_instance +@[simp] lemma monoidal_category.braiding_hom_apply {A B : Rep k G} (x : A) (y : B) : + Action.hom.hom (β_ A B).hom (tensor_product.tmul k x y) = tensor_product.tmul k y x := rfl + +@[simp] lemma monoidal_category.braiding_inv_apply {A B : Rep k G} (x : A) (y : B) : + Action.hom.hom (β_ A B).inv (tensor_product.tmul k y x) = tensor_product.tmul k x y := rfl + section linearization variables (k G) @@ -117,6 +131,56 @@ noncomputable def linearization_of_mul_action_iso (n : ℕ) : (linearization k G).1.1.obj (Action.of_mul_action G (fin n → G)) ≅ of_mul_action k G (fin n → G) := iso.refl _ +variables {k G} + +/-- Given an element `x : A`, there is a natural morphism of representations `k[G] ⟶ A` sending +`g ↦ A.ρ(g)(x).` -/ +@[simps] noncomputable def left_regular_hom (A : Rep k G) (x : A) : + Rep.of_mul_action k G G ⟶ A := +{ hom := finsupp.lift _ _ _ (λ g, A.ρ g x), + comm' := λ g, + begin + refine finsupp.lhom_ext' (λ y, linear_map.ext_ring _), + simpa only [linear_map.comp_apply, Module.comp_def, finsupp.lsingle_apply, + finsupp.lift_apply, Action_ρ_eq_ρ, of_ρ_apply, representation.of_mul_action_single, + finsupp.sum_single_index, zero_smul, one_smul, smul_eq_mul, A.ρ.map_mul], + end } + +lemma left_regular_hom_apply {A : Rep k G} (x : A) : + (left_regular_hom A x).hom (finsupp.single 1 1) = x := +begin + simpa only [left_regular_hom_hom, finsupp.lift_apply, finsupp.sum_single_index, one_smul, + A.ρ.map_one, zero_smul], +end + +/-- Given a `k`-linear `G`-representation `A`, there is a `k`-linear isomorphism between +representation morphisms `Hom(k[G], A)` and `A`. -/ +@[simps] noncomputable def left_regular_hom_equiv (A : Rep k G) : + (Rep.of_mul_action k G G ⟶ A) ≃ₗ[k] A := +{ to_fun := λ f, f.hom (finsupp.single 1 1), + map_add' := λ x y, rfl, + map_smul' := λ r x, rfl, + inv_fun := λ x, left_regular_hom A x, + left_inv := λ f, + begin + refine Action.hom.ext _ _ (finsupp.lhom_ext' (λ (x : G), linear_map.ext_ring _)), + have : f.hom (((of_mul_action k G G).ρ) x (finsupp.single (1 : G) (1 : k))) + = A.ρ x (f.hom (finsupp.single (1 : G) (1 : k))) := + linear_map.ext_iff.1 (f.comm x) (finsupp.single 1 1), + simp only [linear_map.comp_apply, finsupp.lsingle_apply, + left_regular_hom_hom, finsupp.lift_apply, finsupp.sum_single_index, one_smul, ←this, + zero_smul, of_ρ_apply, representation.of_mul_action_single x (1 : G) (1 : k), smul_eq_mul, + mul_one], + end, + right_inv := λ x, left_regular_hom_apply x } + +lemma left_regular_hom_equiv_symm_single {A : Rep k G} (x : A) (g : G) : + ((left_regular_hom_equiv A).symm x).hom (finsupp.single g 1) = A.ρ g x := +begin + simp only [left_regular_hom_equiv_symm_apply, left_regular_hom_hom, + finsupp.lift_apply, finsupp.sum_single_index, zero_smul, one_smul], +end + end linearization end section @@ -162,8 +226,7 @@ rfl `G`-representation `B,` the `k`-linear map underlying the resulting map `B ⟶ iHom(A, A ⊗ B)` is given by flipping the arguments in the natural `k`-bilinear map `A →ₗ[k] B →ₗ[k] A ⊗ B`. -/ @[simp] lemma ihom_coev_app_hom : - Action.hom.hom ((ihom.coev A).app B) = - (tensor_product.mk _ _ _).flip := + Action.hom.hom ((ihom.coev A).app B) = (tensor_product.mk _ _ _).flip := begin refine linear_map.ext (λ x, linear_map.ext (λ y, _)), simpa only [ihom_coev_app_def, functor.map_comp, comp_hom, @@ -208,9 +271,79 @@ the identity map on `Homₖ(A, B).` -/ tensor_product.uncurry _ _ _ _ linear_map.id.flip := monoidal_closed_uncurry_hom _ +variables (A B C) + +/-- There is a `k`-linear isomorphism between the sets of representation morphisms`Hom(A ⊗ B, C)` +and `Hom(B, Homₖ(A, C))`. -/ +noncomputable def monoidal_closed.linear_hom_equiv : + (A ⊗ B ⟶ C) ≃ₗ[k] (B ⟶ (A ⟶[Rep k G] C)) := +{ map_add' := λ f g, rfl, + map_smul' := λ r f, rfl, ..(ihom.adjunction A).hom_equiv _ _ } + +/-- There is a `k`-linear isomorphism between the sets of representation morphisms`Hom(A ⊗ B, C)` +and `Hom(A, Homₖ(B, C))`. -/ +noncomputable def monoidal_closed.linear_hom_equiv_comm : + (A ⊗ B ⟶ C) ≃ₗ[k] (A ⟶ (B ⟶[Rep k G] C)) := +(linear.hom_congr k (β_ A B) (iso.refl _)) ≪≫ₗ + monoidal_closed.linear_hom_equiv _ _ _ + +variables {A B C} + +lemma monoidal_closed.linear_hom_equiv_hom (f : A ⊗ B ⟶ C) : + (monoidal_closed.linear_hom_equiv A B C f).hom = + (tensor_product.curry f.hom).flip := +monoidal_closed_curry_hom _ + +lemma monoidal_closed.linear_hom_equiv_comm_hom (f : A ⊗ B ⟶ C) : + (monoidal_closed.linear_hom_equiv_comm A B C f).hom = + tensor_product.curry f.hom := +begin + dunfold monoidal_closed.linear_hom_equiv_comm, + refine linear_map.ext (λ x, linear_map.ext (λ y, _)), + simp only [linear_equiv.trans_apply, monoidal_closed.linear_hom_equiv_hom, + linear.hom_congr_apply, iso.refl_hom, iso.symm_hom, linear_map.to_fun_eq_coe, + linear_map.coe_comp, function.comp_app, linear.left_comp_apply, linear.right_comp_apply, + category.comp_id, Action.comp_hom, linear_map.flip_apply, tensor_product.curry_apply, + Module.coe_comp, function.comp_app, monoidal_category.braiding_inv_apply], +end + +lemma monoidal_closed.linear_hom_equiv_symm_hom (f : B ⟶ (A ⟶[Rep k G] C)) : + ((monoidal_closed.linear_hom_equiv A B C).symm f).hom = + tensor_product.uncurry k A B C f.hom.flip := +monoidal_closed_uncurry_hom _ + +lemma monoidal_closed.linear_hom_equiv_comm_symm_hom (f : A ⟶ (B ⟶[Rep k G] C)) : + ((monoidal_closed.linear_hom_equiv_comm A B C).symm f).hom = + tensor_product.uncurry k A B C f.hom := +begin + dunfold monoidal_closed.linear_hom_equiv_comm, + refine tensor_product.algebra_tensor_module.curry_injective + (linear_map.ext (λ x, linear_map.ext (λ y, _))), + simp only [linear_equiv.trans_symm, linear_equiv.trans_apply, linear.hom_congr_symm_apply, + iso.refl_inv, linear_map.coe_comp, function.comp_app, category.comp_id, Action.comp_hom, + monoidal_closed.linear_hom_equiv_symm_hom, tensor_product.algebra_tensor_module.curry_apply, + linear_map.coe_restrict_scalars, linear_map.to_fun_eq_coe, linear_map.flip_apply, + tensor_product.curry_apply, Module.coe_comp, function.comp_app, + monoidal_category.braiding_hom_apply, tensor_product.uncurry_apply], +end + end end Rep +namespace representation +variables {k G : Type u} [comm_ring k] [monoid G] {V W : Type u} + [add_comm_group V] [add_comm_group W] [module k V] [module k W] + (ρ : representation k G V) (τ : representation k G W) + +/-- Tautological isomorphism to help Lean in typechecking. -/ +def Rep_of_tprod_iso : Rep.of (ρ.tprod τ) ≅ Rep.of ρ ⊗ Rep.of τ := iso.refl _ + +lemma Rep_of_tprod_iso_apply (x : tensor_product k V W) : + (Rep_of_tprod_iso ρ τ).hom.hom x = x := rfl + +lemma Rep_of_tprod_iso_inv_apply (x : tensor_product k V W) : + (Rep_of_tprod_iso ρ τ).inv.hom x = x := rfl +end representation /-! # The categorical equivalence `Rep k G ≌ Module.{u} (monoid_algebra k G)`. -/ diff --git a/src/representation_theory/basic.lean b/src/representation_theory/basic.lean index e04c5c5008466..2db82fe660fdd 100644 --- a/src/representation_theory/basic.lean +++ b/src/representation_theory/basic.lean @@ -245,6 +245,10 @@ variables {k G H} lemma of_mul_action_def (g : G) : of_mul_action k G H g = finsupp.lmap_domain k k ((•) g) := rfl +lemma of_mul_action_single (g : G) (x : H) (r : k) : + of_mul_action k G H g (finsupp.single x r) = finsupp.single (g • x) r := +finsupp.map_domain_single + end mul_action section group diff --git a/src/representation_theory/group_cohomology_resolution.lean b/src/representation_theory/group_cohomology_resolution.lean index 081d7f5293541..987be7a8fe044 100644 --- a/src/representation_theory/group_cohomology_resolution.lean +++ b/src/representation_theory/group_cohomology_resolution.lean @@ -169,12 +169,12 @@ variables {k G n} @[simp] lemma to_tensor_single (f : Gⁿ⁺¹) (m : k) : (to_tensor k G n).hom (single f m) = single (f 0) m ⊗ₜ single (λ i, (f i)⁻¹ * f i.succ) 1 := -to_tensor_aux_single _ _ +by apply to_tensor_aux_single f m @[simp] lemma of_tensor_single (g : G) (m : k) (x : Gⁿ →₀ k) : (of_tensor k G n).hom ((single g m) ⊗ₜ x) = finsupp.lift (Rep.of_mul_action k G Gⁿ⁺¹) k Gⁿ (λ f, single (g • partial_prod f) m) x := -of_tensor_aux_single _ _ _ +by apply of_tensor_aux_single g m x lemma of_tensor_single' (g : G →₀ k) (x : Gⁿ) (m : k) : (of_tensor k G n).hom (g ⊗ₜ single x m) = @@ -233,6 +233,58 @@ module.free.of_basis (of_mul_action_basis k G n) end basis end group_cohomology.resolution +namespace Rep +variables (n) [group G] +open group_cohomology.resolution + +/-- Given a `k`-linear `G`-representation `A`, the set of representation morphisms +`Hom(k[Gⁿ⁺¹], A)` is `k`-linearly isomorphic to the set of functions `Gⁿ → A`. -/ +noncomputable def diagonal_hom_equiv (A : Rep k G) : + (Rep.of_mul_action k G (fin (n + 1) → G) ⟶ A) ≃ₗ[k] ((fin n → G) → A) := +linear.hom_congr k ((equiv_tensor k G n).trans + ((representation.of_mul_action k G G).Rep_of_tprod_iso 1)) (iso.refl _) ≪≫ₗ + ((Rep.monoidal_closed.linear_hom_equiv_comm _ _ _) ≪≫ₗ (Rep.left_regular_hom_equiv _)) + ≪≫ₗ (finsupp.llift A k k (fin n → G)).symm + +variables {n} + +/-- Given a `k`-linear `G`-representation `A`, `diagonal_hom_equiv` is a `k`-linear isomorphism of +the set of representation morphisms `Hom(k[Gⁿ⁺¹], A)` with `Fun(Gⁿ, A)`. This lemma says that this +sends a morphism of representations `f : k[Gⁿ⁺¹] ⟶ A` to the function +`(g₁, ..., gₙ) ↦ f(1, g₁, g₁g₂, ..., g₁g₂...gₙ).` -/ +lemma diagonal_hom_equiv_apply {A : Rep k G} (f : Rep.of_mul_action k G (fin (n + 1) → G) ⟶ A) + (x : fin n → G) : diagonal_hom_equiv n A f x = f.hom (finsupp.single (fin.partial_prod x) 1) := +begin + unfold diagonal_hom_equiv, + simpa only [linear_equiv.trans_apply, Rep.left_regular_hom_equiv_apply, + monoidal_closed.linear_hom_equiv_comm_hom, finsupp.llift_symm_apply, tensor_product.curry_apply, + linear.hom_congr_apply, iso.refl_hom, iso.trans_inv, Action.comp_hom, Module.comp_def, + linear_map.comp_apply, equiv_tensor_inv_def, representation.Rep_of_tprod_iso_inv_apply, + of_tensor_single (1 : G) (1 : k) (finsupp.single x (1 : k)), finsupp.lift_apply, + finsupp.sum_single_index, one_smul, zero_smul], +end + +/-- Given a `k`-linear `G`-representation `A`, `diagonal_hom_equiv` is a `k`-linear isomorphism of +the set of representation morphisms `Hom(k[Gⁿ⁺¹], A)` with `Fun(Gⁿ, A)`. This lemma says that the +inverse map sends a function `f : Gⁿ → A` to the representation morphism sending +`(g₀, ... gₙ) ↦ ρ(g₀)(f(g₀⁻¹g₁, g₁⁻¹g₂, ..., gₙ₋₁⁻¹gₙ))`, where `ρ` is the representation attached +to `A`. -/ +lemma diagonal_hom_equiv_symm_apply {A : Rep k G} (f : (fin n → G) → A) (x : fin (n + 1) → G) : + ((diagonal_hom_equiv n A).symm f).hom (finsupp.single x 1) + = A.ρ (x 0) (f (λ (i : fin n), (x ↑i)⁻¹ * x i.succ)) := +begin + unfold diagonal_hom_equiv, + simp only [linear_equiv.trans_symm, linear_equiv.symm_symm, linear_equiv.trans_apply, + Rep.left_regular_hom_equiv_symm_apply, linear.hom_congr_symm_apply, Action.comp_hom, + iso.refl_inv, category.comp_id, Rep.monoidal_closed.linear_hom_equiv_comm_symm_hom, + iso.trans_hom, Module.comp_def, linear_map.comp_apply, representation.Rep_of_tprod_iso_apply, + equiv_tensor_def, to_tensor_single x (1 : k), tensor_product.uncurry_apply, + Rep.left_regular_hom_hom, finsupp.lift_apply, Rep.ihom_obj_ρ, representation.lin_hom_apply, + finsupp.sum_single_index, zero_smul, one_smul, Rep.of_ρ, + monoid_hom.one_apply, linear_map.one_apply, finsupp.llift_apply A k k], +end + +end Rep variables (G) /-- The simplicial `G`-set sending `[n]` to `Gⁿ⁺¹` equipped with the diagonal action of `G`. -/ From f209a5a913a8aded802138f3aa82a192fa8e3697 Mon Sep 17 00:00:00 2001 From: Jakob von Raumer Date: Mon, 27 Mar 2023 08:37:39 +0000 Subject: [PATCH 0700/1291] feat(data/zmod/basic): `zmod.val_min_abs` lemmas (#18119) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Yaël Dillies --- src/data/zmod/basic.lean | 28 ++++++++++++++++++++++++++++ 1 file changed, 28 insertions(+) diff --git a/src/data/zmod/basic.lean b/src/data/zmod/basic.lean index e6472d0723fd0..f751ced32c8f4 100644 --- a/src/data/zmod/basic.lean +++ b/src/data/zmod/basic.lean @@ -33,6 +33,8 @@ This is a ring hom if the ring has characteristic dividing `n` -/ +open function + namespace zmod instance : char_zero (zmod 0) := (by apply_instance : char_zero ℤ) @@ -900,6 +902,8 @@ lemma prime_ne_zero (p q : ℕ) [hp : fact p.prime] [hq : fact q.prime] (hpq : p by rwa [← nat.cast_zero, ne.def, eq_iff_modeq_nat, nat.modeq_zero_iff_dvd, ← hp.1.coprime_iff_not_dvd, nat.coprime_primes hp.1 hq.1] +variables {n a : ℕ} + lemma val_min_abs_nat_abs_eq_min {n : ℕ} [hpos : ne_zero n] (a : zmod n) : a.val_min_abs.nat_abs = min a.val (n - a.val) := begin @@ -916,6 +920,30 @@ begin apply nat.lt_succ_self } end +lemma val_min_abs_nat_cast_of_le_half (ha : a ≤ n / 2) : (a : zmod n).val_min_abs = a := +begin + cases n, + { simp }, + { simp [val_min_abs_def_pos, val_nat_cast, + nat.mod_eq_of_lt (ha.trans_lt $ nat.div_lt_self' _ 0), ha] } +end + +lemma val_min_abs_nat_cast_of_half_lt (ha : n / 2 < a) (ha' : a < n) : + (a : zmod n).val_min_abs = a - n := +begin + cases n, + { cases not_lt_bot ha' }, + { simp [val_min_abs_def_pos, val_nat_cast, nat.mod_eq_of_lt ha', ha.not_le] } +end + +@[simp] lemma val_min_nat_abs_nat_cast_eq_self [ne_zero n] : + (a : zmod n).val_min_abs = a ↔ a ≤ n / 2 := +begin + refine ⟨λ ha, _, val_min_abs_nat_cast_of_le_half⟩, + rw [←int.nat_abs_of_nat a, ←ha], + exact nat_abs_val_min_abs_le a, +end + lemma nat_abs_min_of_le_div_two (n : ℕ) (x y : ℤ) (he : (x : zmod n) = y) (hl : x.nat_abs ≤ n / 2) : x.nat_abs ≤ y.nat_abs := begin From b68f3403d74afcfd65ab3fe65220572b410a3ed4 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 27 Mar 2023 10:14:39 +0000 Subject: [PATCH 0701/1291] feat(ring_theory/trace): trace of product rings (#18656) Probably not a very useful result, but thought I'd add it to get some intuition for `trace`. I didn't bother with `pi` because we don't have the analogous `linear_map.trace` statements either. --- src/ring_theory/trace.lean | 18 ++++++++++++++++++ 1 file changed, 18 insertions(+) diff --git a/src/ring_theory/trace.lean b/src/ring_theory/trace.lean index 91403bc5ddab4..10188e27f7fe2 100644 --- a/src/ring_theory/trace.lean +++ b/src/ring_theory/trace.lean @@ -161,6 +161,24 @@ lemma trace_comp_trace [algebra K T] [algebra L T] [is_scalar_tower K L T] (trace K L).comp ((trace L T).restrict_scalars K) = trace K T := by { ext, rw [linear_map.comp_apply, linear_map.restrict_scalars_apply, trace_trace] } +@[simp] +lemma trace_prod_apply + [module.free R S] [module.free R T] [module.finite R S] [module.finite R T] + (x : S × T) : trace R (S × T) x = trace R S x.fst + trace R T x.snd := +begin + nontriviality R, + let f := (lmul R S).to_linear_map.prod_map (lmul R T).to_linear_map, + have : (lmul R (S × T)).to_linear_map = (prod_map_linear R S T S T R).comp f := + linear_map.ext₂ prod.mul_def, + simp_rw [trace, this], + exact trace_prod_map' _ _, +end + +lemma trace_prod + [module.free R S] [module.free R T] [module.finite R S] [module.finite R T] : + trace R (S × T) = (trace R S).coprod (trace R T) := +linear_map.ext $ λ p, by rw [coprod_apply, trace_prod_apply] + section trace_form variables (R S) From 9015c511549dc77a0f8d6eba021d8ac4bba20c82 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 27 Mar 2023 10:14:40 +0000 Subject: [PATCH 0702/1291] feat(ring_theory/complex): trace and norm of a complex number (#18658) --- src/ring_theory/complex.lean | 43 ++++++++++++++++++++++++++++++++++++ 1 file changed, 43 insertions(+) create mode 100644 src/ring_theory/complex.lean diff --git a/src/ring_theory/complex.lean b/src/ring_theory/complex.lean new file mode 100644 index 0000000000000..598d189f045c2 --- /dev/null +++ b/src/ring_theory/complex.lean @@ -0,0 +1,43 @@ +/- +Copyright (c) 2023 Eric Wieser. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Eric Wieser +-/ +import data.complex.module +import ring_theory.norm +import ring_theory.trace + +/-! # Lemmas about `algebra.trace` and `algebra.norm` on `ℂ` -/ + +open complex + +lemma algebra.left_mul_matrix_complex (z : ℂ) : + algebra.left_mul_matrix complex.basis_one_I z = !![z.re, -z.im; z.im, z.re] := +begin + ext i j, + rw [algebra.left_mul_matrix_eq_repr_mul, complex.coe_basis_one_I_repr, complex.coe_basis_one_I, + mul_re, mul_im, matrix.of_apply], + fin_cases j, + { simp_rw [matrix.cons_val_zero, one_re, one_im, mul_zero, mul_one, sub_zero, zero_add], + fin_cases i; refl }, + { simp_rw [matrix.cons_val_one, matrix.head_cons, I_re, I_im, mul_zero, mul_one, zero_sub, + add_zero], + fin_cases i; refl }, +end + +lemma algebra.trace_complex_apply (z : ℂ) : algebra.trace ℝ ℂ z = 2*z.re := +begin + rw [algebra.trace_eq_matrix_trace complex.basis_one_I, + algebra.left_mul_matrix_complex, matrix.trace_fin_two], + exact (two_mul _).symm +end + +lemma algebra.norm_complex_apply (z : ℂ) : algebra.norm ℝ z = z.norm_sq := +begin + rw [algebra.norm_eq_matrix_det complex.basis_one_I, + algebra.left_mul_matrix_complex, matrix.det_fin_two, norm_sq_apply], + simp, +end + +lemma algebra.norm_complex_eq : algebra.norm ℝ = norm_sq.to_monoid_hom := +monoid_hom.ext algebra.norm_complex_apply From b8627dbac120a9ad6267a75575ae1e070d5bff5b Mon Sep 17 00:00:00 2001 From: Moritz Doll Date: Mon, 27 Mar 2023 11:20:17 +0000 Subject: [PATCH 0703/1291] refactor(topology/algebra/module/strong_topology): split of local convexity (#18671) The reason for this split is not only to reduce the import tree, but also to find a good home for proving `with_seminorm` versions of the local convexity results. --- .../locally_convex/strong_topology.lean | 67 +++++++++++++++++++ .../algebra/module/strong_topology.lean | 22 ------ 2 files changed, 67 insertions(+), 22 deletions(-) create mode 100644 src/analysis/locally_convex/strong_topology.lean diff --git a/src/analysis/locally_convex/strong_topology.lean b/src/analysis/locally_convex/strong_topology.lean new file mode 100644 index 0000000000000..79306e7a3b0ac --- /dev/null +++ b/src/analysis/locally_convex/strong_topology.lean @@ -0,0 +1,67 @@ +/- +Copyright (c) 2022 Anatole Dedecker. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Anatole Dedecker +-/ +import topology.algebra.module.strong_topology +import topology.algebra.module.locally_convex + +/-! +# Local convexity of the strong topology + +In this file we prove that the strong topology on `E →L[ℝ] F` is locally convex provided that `F` is +locally convex. + +## References + +* [N. Bourbaki, *Topological Vector Spaces*][bourbaki1987] + +## Todo + +* Characterization in terms of seminorms + +## Tags + +locally convex, bounded convergence +-/ + +open_locale topology uniform_convergence + +variables {E F : Type*} + +namespace continuous_linear_map + +section general + +variables [add_comm_group E] [module ℝ E] [topological_space E] + [add_comm_group F] [module ℝ F] [topological_space F] [topological_add_group F] + [has_continuous_const_smul ℝ F] [locally_convex_space ℝ F] + +lemma strong_topology.locally_convex_space (𝔖 : set (set E)) (h𝔖₁ : 𝔖.nonempty) + (h𝔖₂ : directed_on (⊆) 𝔖) : + @locally_convex_space ℝ (E →L[ℝ] F) _ _ _ (strong_topology (ring_hom.id ℝ) F 𝔖) := +begin + letI : topological_space (E →L[ℝ] F) := strong_topology (ring_hom.id ℝ) F 𝔖, + haveI : topological_add_group (E →L[ℝ] F) := strong_topology.topological_add_group _ _ _, + refine locally_convex_space.of_basis_zero _ _ _ _ + (strong_topology.has_basis_nhds_zero_of_basis _ _ _ h𝔖₁ h𝔖₂ + (locally_convex_space.convex_basis_zero ℝ F)) _, + rintros ⟨S, V⟩ ⟨hS, hVmem, hVconvex⟩ f hf g hg a b ha hb hab x hx, + exact hVconvex (hf x hx) (hg x hx) ha hb hab, +end + +end general + +section bounded_sets + +variables [add_comm_group E] [module ℝ E] [topological_space E] + [add_comm_group F] [module ℝ F] [topological_space F] [topological_add_group F] + [has_continuous_const_smul ℝ F] [locally_convex_space ℝ F] + +instance : locally_convex_space ℝ (E →L[ℝ] F) := +strong_topology.locally_convex_space _ ⟨∅, bornology.is_vonN_bounded_empty ℝ E⟩ + (directed_on_of_sup_mem $ λ _ _, bornology.is_vonN_bounded.union) + +end bounded_sets + +end continuous_linear_map diff --git a/src/topology/algebra/module/strong_topology.lean b/src/topology/algebra/module/strong_topology.lean index 880e183841291..9b60004ef1c63 100644 --- a/src/topology/algebra/module/strong_topology.lean +++ b/src/topology/algebra/module/strong_topology.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Anatole Dedecker -/ import topology.algebra.uniform_convergence -import topology.algebra.module.locally_convex /-! # Strong topologies on the space of continuous linear maps @@ -47,7 +46,6 @@ sets). ## TODO -* show that these topologies are T₂ and locally convex if the topology on `F` is * add a type alias for continuous linear maps with the topology of `𝔖`-convergence? ## Tags @@ -170,20 +168,6 @@ lemma strong_topology.has_basis_nhds_zero [topological_space F] [topological_add (λ SV, {f : E →SL[σ] F | ∀ x ∈ SV.1, f x ∈ SV.2}) := strong_topology.has_basis_nhds_zero_of_basis σ F 𝔖 h𝔖₁ h𝔖₂ (𝓝 0).basis_sets -lemma strong_topology.locally_convex_space [topological_space F'] - [topological_add_group F'] [has_continuous_const_smul ℝ F'] [locally_convex_space ℝ F'] - (𝔖 : set (set E')) (h𝔖₁ : 𝔖.nonempty) (h𝔖₂ : directed_on (⊆) 𝔖) : - @locally_convex_space ℝ (E' →L[ℝ] F') _ _ _ (strong_topology (ring_hom.id ℝ) F' 𝔖) := -begin - letI : topological_space (E' →L[ℝ] F') := strong_topology (ring_hom.id ℝ) F' 𝔖, - haveI : topological_add_group (E' →L[ℝ] F') := strong_topology.topological_add_group _ _ _, - refine locally_convex_space.of_basis_zero _ _ _ _ - (strong_topology.has_basis_nhds_zero_of_basis _ _ _ h𝔖₁ h𝔖₂ - (locally_convex_space.convex_basis_zero ℝ F')) _, - rintros ⟨S, V⟩ ⟨hS, hVmem, hVconvex⟩ f hf g hg a b ha hb hab x hx, - exact hVconvex (hf x hx) (hg x hx) ha hb hab, -end - end general section bounded_sets @@ -237,12 +221,6 @@ protected lemma has_basis_nhds_zero [topological_space F] (λ SV, {f : E →SL[σ] F | ∀ x ∈ SV.1, f x ∈ SV.2}) := continuous_linear_map.has_basis_nhds_zero_of_basis (𝓝 0).basis_sets -instance [topological_space E'] [topological_space F'] [topological_add_group F'] - [has_continuous_const_smul ℝ F'] [locally_convex_space ℝ F'] : - locally_convex_space ℝ (E' →L[ℝ] F') := -strong_topology.locally_convex_space _ ⟨∅, bornology.is_vonN_bounded_empty ℝ E'⟩ - (directed_on_of_sup_mem $ λ _ _, bornology.is_vonN_bounded.union) - end bounded_sets end continuous_linear_map From 728baa2f54e6062c5879a3e397ac6bac323e506f Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Mon, 27 Mar 2023 13:45:21 +0000 Subject: [PATCH 0704/1291] feat(archive/imo/imo2006_q5): IMO 2006 Q5 (#15613) See module docstring for a thorough explanation of the proof. Co-authored-by: Thomas Browning --- archive/imo/imo2006_q5.lean | 210 +++++++++++++++++++++++++ src/data/int/basic.lean | 3 + src/data/int/order/basic.lean | 10 ++ src/data/list/cycle.lean | 4 + src/data/polynomial/degree/lemmas.lean | 8 + src/data/polynomial/eval.lean | 14 ++ 6 files changed, 249 insertions(+) create mode 100644 archive/imo/imo2006_q5.lean diff --git a/archive/imo/imo2006_q5.lean b/archive/imo/imo2006_q5.lean new file mode 100644 index 0000000000000..665882fd84f23 --- /dev/null +++ b/archive/imo/imo2006_q5.lean @@ -0,0 +1,210 @@ +/- +Copyright (c) 2022 Violeta Hernández Palacios. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Violeta Hernández Palacios +-/ + +import data.polynomial.ring_division +import dynamics.periodic_pts + +/-! +# IMO 2006 Q5 + +Let $P(x)$ be a polynomial of degree $n>1$ with integer coefficients, and let $k$ be a positive +integer. Consider the polynomial $Q(x) = P(P(\ldots P(P(x))\ldots))$, where $P$ occurs $k$ times. +Prove that there are at most $n$ integers $t$ such that $Q(t)=t$. + +## Sketch of solution + +The following solution is adapted from +https://artofproblemsolving.com/wiki/index.php/2006_IMO_Problems/Problem_5. + +Let $P^k$ denote the polynomial $P$ composed with itself $k$ times. We rely on a key observation: if +$P^k(t)=t$, then $P(P(t))=t$. We prove this by building the cyclic list +$(P(t)-t,P^2(t)-P(t),\ldots)$, and showing that each entry divides the next, which by transitivity +implies they all divide each other, and thus have the same absolute value. + +If the entries in this list are all pairwise equal, then we can show inductively that for positive +$n$, $P^n(t)-t$ must always have the same sign as $P(t)-t$. Substituting $n=k$ gives us $P(t)=t$ and +in particular $P(P(t))=t$. + +Otherwise, there must be two consecutive entries that are opposites of one another. This means +$P^{n+2}(t)-P^{n+1}(t)=P^n(t)-P^{n+1}(t)$, which implies $P^{n+2}(t)=P^n(t)$ and $P(P(t))=t$. + +With this lemma, we can reduce the problem to the case $k=2$. If every root of $P(P(t))-t$ is also a +root of $P(t)-t$, then we're done. Otherwise, there exist $a$ and $b$ with $a\ne b$ and $P(a)=b$, +$P(b)=a$. For any root $t$ of $P(P(t))-t$, defining $u=P(t)$, we easily verify $a-t\mid b-u$, +$b-u\mid a-t$, $a-u\mid b-t$, $b-t\mid a-u$, which imply $|a-t|=|b-u|$ and $|a-u|=|b-t|$. By casing +on these equalities, we deduce $a+b=t+u$. This means that every root of $P(P(t))-t$ is a root of +$P(t)+t-a-b$, and we're again done. +-/ + +open function polynomial + +/-- If every entry in a cyclic list of integers divides the next, then they all have the same +absolute value. -/ +theorem int.nat_abs_eq_of_chain_dvd {l : cycle ℤ} {x y : ℤ} (hl : l.chain (∣)) + (hx : x ∈ l) (hy : y ∈ l) : x.nat_abs = y.nat_abs := +begin + rw cycle.chain_iff_pairwise at hl, + exact int.nat_abs_eq_of_dvd_dvd (hl x hx y hy) (hl y hy x hx) +end + +theorem int.add_eq_add_of_nat_abs_eq_of_nat_abs_eq {a b c d : ℤ} (hne : a ≠ b) + (h₁ : (c - a).nat_abs = (d - b).nat_abs) (h₂ : (c - b).nat_abs = (d - a).nat_abs) : + a + b = c + d := +begin + cases int.nat_abs_eq_nat_abs_iff.1 h₁ with h₁ h₁, + { cases int.nat_abs_eq_nat_abs_iff.1 h₂ with h₂ h₂, + { exact (hne $ by linarith).elim }, + { linarith } }, + { linarith } +end + +/-- The main lemma in the proof: if $P^k(t)=t$, then $P(P(t))=t$. -/ +theorem polynomial.is_periodic_pt_eval_two {P : polynomial ℤ} {t : ℤ} + (ht : t ∈ periodic_pts (λ x, P.eval x)) : is_periodic_pt (λ x, P.eval x) 2 t := +begin + -- The cycle [P(t) - t, P(P(t)) - P(t), ...] + let C : cycle ℤ := (periodic_orbit (λ x, P.eval x) t).map (λ x, P.eval x - x), + have HC : ∀ {n : ℕ}, (λ x, P.eval x)^[n + 1] t - ((λ x, P.eval x)^[n] t) ∈ C, + { intro n, + rw [cycle.mem_map, function.iterate_succ_apply'], + exact ⟨_, iterate_mem_periodic_orbit ht n, rfl⟩ }, + + -- Elements in C are all divisible by one another. + have Hdvd : C.chain (∣), + { rw [cycle.chain_map, periodic_orbit_chain' _ ht], + intro n, + convert sub_dvd_eval_sub ((λ x, P.eval x)^[n + 1] t) ((λ x, P.eval x)^[n] t) P; + rw function.iterate_succ_apply' }, + + -- Any two entries in C have the same absolute value. + have Habs : ∀ m n : ℕ, ((λ x, P.eval x)^[m + 1] t - ((λ x, P.eval x)^[m] t)).nat_abs = + ((λ x, P.eval x)^[n + 1] t - ((λ x, P.eval x)^[n] t)).nat_abs := + λ m n, int.nat_abs_eq_of_chain_dvd Hdvd HC HC, + + -- We case on whether the elements on C are pairwise equal. + by_cases HC' : C.chain (=), + { -- Any two entries in C are equal. + have Heq : ∀ m n : ℕ, (λ x, P.eval x)^[m + 1] t - ((λ x, P.eval x)^[m] t) = + ((λ x, P.eval x)^[n + 1] t - ((λ x, P.eval x)^[n] t)) := + λ m n, cycle.chain_iff_pairwise.1 HC' _ HC _ HC, + + -- The sign of P^n(t) - t is the same as P(t) - t for positive n. Proven by induction on n. + have IH : ∀ n : ℕ, ((λ x, P.eval x)^[n + 1] t - t).sign = (P.eval t - t).sign, + { intro n, + induction n with n IH, + { refl }, + { apply eq.trans _ (int.sign_add_eq_of_sign_eq IH), + have H := Heq n.succ 0, + dsimp at H ⊢, + rw [←H, sub_add_sub_cancel'] } }, + + -- This implies that the sign of P(t) - t is the same as the sign of P^k(t) - t, which is 0. + -- Hence P(t) = t and P(P(t)) = P(t). + rcases ht with ⟨(_ | k), hk, hk'⟩, + { exact (irrefl 0 hk).elim }, + { have H := IH k, + rw [hk'.is_fixed_pt.eq, sub_self, int.sign_zero, eq_comm, int.sign_eq_zero_iff_zero, + sub_eq_zero] at H, + simp [is_periodic_pt, is_fixed_pt, H] } }, + { -- We take two nonequal consecutive entries. + rw [cycle.chain_map, periodic_orbit_chain' _ ht] at HC', + push_neg at HC', + cases HC' with n hn, + + -- They must have opposite sign, so that P^{k + 1}(t) - P^k(t) = P^{k + 2}(t) - P^{k + 1}(t). + cases int.nat_abs_eq_nat_abs_iff.1 (Habs n n.succ) with hn' hn', + { apply (hn _).elim, + convert hn'; + simp only [function.iterate_succ_apply'] }, + + -- We deduce P^{k + 2}(t) = P^k(t) and hence P(P(t)) = t. + { rw [neg_sub, sub_right_inj] at hn', + simp only [function.iterate_succ_apply'] at hn', + exact @is_periodic_pt_of_mem_periodic_pts_of_is_periodic_pt_iterate _ _ t 2 n ht hn'.symm } } +end + +theorem polynomial.iterate_comp_sub_X_ne {P : polynomial ℤ} (hP : 1 < P.nat_degree) {k : ℕ} + (hk : 0 < k) : P.comp^[k] X - X ≠ 0 := +by { rw sub_ne_zero, apply_fun nat_degree, simpa using (one_lt_pow hP hk.ne').ne' } + +/-- We solve the problem for the specific case k = 2 first. -/ +theorem imo2006_q5' {P : polynomial ℤ} (hP : 1 < P.nat_degree) : + (P.comp P - X).roots.to_finset.card ≤ P.nat_degree := +begin + -- Auxiliary lemmas on degrees. + have hPX : (P - X).nat_degree = P.nat_degree, + { rw nat_degree_sub_eq_left_of_nat_degree_lt, + simpa using hP }, + have hPX' : P - X ≠ 0, + { intro h, + rw [h, nat_degree_zero] at hPX, + rw ←hPX at hP, + exact (zero_le_one.not_lt hP).elim }, + + -- If every root of P(P(t)) - t is also a root of P(t) - t, then we're done. + by_cases H : (P.comp P - X).roots.to_finset ⊆ (P - X).roots.to_finset, + { exact (finset.card_le_of_subset H).trans ((multiset.to_finset_card_le _).trans + ((card_roots' _).trans_eq hPX)) }, + + -- Otherwise, take a, b with P(a) = b, P(b) = a, a ≠ b. + { rcases finset.not_subset.1 H with ⟨a, ha, hab⟩, + replace ha := is_root_of_mem_roots (multiset.mem_to_finset.1 ha), + simp [sub_eq_zero] at ha, + simp [mem_roots hPX'] at hab, + set b := P.eval a, + rw sub_eq_zero at hab, + + -- More auxiliary lemmas on degrees. + have hPab : (P + X - a - b).nat_degree = P.nat_degree, + { rw [sub_sub, ←int.cast_add], + have h₁ : (P + X).nat_degree = P.nat_degree, + { rw nat_degree_add_eq_left_of_nat_degree_lt, + simpa using hP }, + rw nat_degree_sub_eq_left_of_nat_degree_lt; + rwa h₁, + rw nat_degree_int_cast, + exact zero_lt_one.trans hP }, + have hPab' : P + X - a - b ≠ 0, + { intro h, + rw [h, nat_degree_zero] at hPab, + rw ←hPab at hP, + exact (zero_le_one.not_lt hP).elim }, + + -- We claim that every root of P(P(t)) - t is a root of P(t) + t - a - b. This allows us to + -- conclude the problem. + suffices H' : (P.comp P - X).roots.to_finset ⊆ (P + X - a - b).roots.to_finset, + { exact (finset.card_le_of_subset H').trans ((multiset.to_finset_card_le _).trans $ + (card_roots' _).trans_eq hPab) }, + + { -- Let t be a root of P(P(t)) - t, define u = P(t). + intros t ht, + replace ht := is_root_of_mem_roots (multiset.mem_to_finset.1 ht), + simp [sub_eq_zero] at ht, + simp only [mem_roots hPab', sub_eq_iff_eq_add, multiset.mem_to_finset, is_root.def, eval_sub, + eval_add, eval_X, eval_C, eval_int_cast, int.cast_id, zero_add], + + -- An auxiliary lemma proved earlier implies we only need to show |t - a| = |u - b| and + -- |t - b| = |u - a|. We prove this by establishing that each side of either equation divides + -- the other. + apply (int.add_eq_add_of_nat_abs_eq_of_nat_abs_eq hab _ _).symm; + apply int.nat_abs_eq_of_dvd_dvd; + set u := P.eval t, + { rw [←ha, ←ht], apply sub_dvd_eval_sub }, + { apply sub_dvd_eval_sub }, + { rw ←ht, apply sub_dvd_eval_sub }, + { rw ←ha, apply sub_dvd_eval_sub } } } +end + +/-- The general problem follows easily from the k = 2 case. -/ +theorem imo2006_q5 {P : polynomial ℤ} (hP : 1 < P.nat_degree) {k : ℕ} (hk : 0 < k) : + (P.comp^[k] X - X).roots.to_finset.card ≤ P.nat_degree := +begin + apply (finset.card_le_of_subset $ λ t ht, _).trans (imo2006_q5' hP), + have hP' : P.comp P - X ≠ 0 := by simpa using polynomial.iterate_comp_sub_X_ne hP zero_lt_two, + replace ht := is_root_of_mem_roots (multiset.mem_to_finset.1 ht), + simp only [sub_eq_zero, is_root.def, eval_sub, iterate_comp_eval, eval_X] at ht, + simpa [mem_roots hP', sub_eq_zero] using polynomial.is_periodic_pt_eval_two ⟨k, hk, ht⟩ +end diff --git a/src/data/int/basic.lean b/src/data/int/basic.lean index 1554d6cc2650e..d780c83e2a63d 100644 --- a/src/data/int/basic.lean +++ b/src/data/int/basic.lean @@ -86,6 +86,9 @@ namespace int @[simp] lemma add_neg_one (i : ℤ) : i + -1 = i - 1 := rfl +@[simp] theorem sign_coe_add_one (n : ℕ) : int.sign (n + 1) = 1 := rfl +@[simp] theorem sign_neg_succ_of_nat (n : ℕ) : int.sign -[1+ n] = -1 := rfl + @[simp] lemma default_eq_zero : default = (0 : ℤ) := rfl meta instance : has_to_format ℤ := ⟨λ z, to_string z⟩ diff --git a/src/data/int/order/basic.lean b/src/data/int/order/basic.lean index 05db2a22bca2c..4f84a3537d7aa 100644 --- a/src/data/int/order/basic.lean +++ b/src/data/int/order/basic.lean @@ -76,6 +76,16 @@ lemma coe_nat_ne_zero_iff_pos {n : ℕ} : (n : ℤ) ≠ 0 ↔ 0 < n := @[norm_cast] lemma abs_coe_nat (n : ℕ) : |(n : ℤ)| = n := abs_of_nonneg (coe_nat_nonneg n) +theorem sign_add_eq_of_sign_eq : ∀ {m n : ℤ}, m.sign = n.sign → (m + n).sign = n.sign := +begin + have : (1 : ℤ) ≠ -1 := dec_trivial, + rintro ((_ | m) | m) ((_ | n) | n); + simp [this, this.symm], + rw int.sign_eq_one_iff_pos, + apply int.add_pos; + { exact zero_lt_one.trans_le (le_add_of_nonneg_left $ coe_zero_le _) } +end + /-! ### succ and pred -/ theorem lt_succ_self (a : ℤ) : a < succ a := diff --git a/src/data/list/cycle.lean b/src/data/list/cycle.lean index 5debeb8266bc0..173833768ea84 100644 --- a/src/data/list/cycle.lean +++ b/src/data/list/cycle.lean @@ -626,6 +626,10 @@ rfl @[simp] lemma map_eq_nil {β : Type*} (f : α → β) (s : cycle α) : map f s = nil ↔ s = nil := quotient.induction_on' s (by simp) +@[simp] lemma mem_map {β : Type*} {f : α → β} {b : β} {s : cycle α} : + b ∈ s.map f ↔ ∃ a, a ∈ s ∧ f a = b := +quotient.induction_on' s (by simp) + /-- The `multiset` of lists that can make the cycle. -/ def lists (s : cycle α) : multiset (list α) := quotient.lift_on' s diff --git a/src/data/polynomial/degree/lemmas.lean b/src/data/polynomial/degree/lemmas.lean index 43ca73a90d635..260b20daee6a5 100644 --- a/src/data/polynomial/degree/lemmas.lean +++ b/src/data/polynomial/degree/lemmas.lean @@ -327,6 +327,14 @@ begin ne_zero_of_nat_degree_gt (nat.pos_of_ne_zero q0), pow_ne_zero, ne.def, not_false_iff] } end +@[simp] theorem nat_degree_iterate_comp (k : ℕ) : + (p.comp^[k] q).nat_degree = p.nat_degree ^ k * q.nat_degree := +begin + induction k with k IH, + { simp }, + { rw [function.iterate_succ_apply', nat_degree_comp, IH, pow_succ, mul_assoc] } +end + lemma leading_coeff_comp (hq : nat_degree q ≠ 0) : leading_coeff (p.comp q) = leading_coeff p * leading_coeff q ^ nat_degree p := by rw [← coeff_comp_degree_mul_degree hq, ← nat_degree_comp, coeff_nat_degree] diff --git a/src/data/polynomial/eval.lean b/src/data/polynomial/eval.lean index b57008b992ac0..c583a720f61b6 100644 --- a/src/data/polynomial/eval.lean +++ b/src/data/polynomial/eval.lean @@ -783,6 +783,15 @@ lemma eval₂_comp {x : S} : eval₂ f x (p.comp q) = eval₂ f (eval₂ f x q) p := by rw [comp, p.as_sum_range]; simp [eval₂_finset_sum, eval₂_pow] +@[simp] +lemma iterate_comp_eval₂ (k : ℕ) (t : S) : + eval₂ f t (p.comp^[k] q) = ((λ x, eval₂ f x p)^[k] (eval₂ f t q)) := +begin + induction k with k IH, + { simp }, + { rw [function.iterate_succ_apply', function.iterate_succ_apply', eval₂_comp, IH] } +end + end section @@ -808,6 +817,11 @@ begin { intros n a, simp, } end +@[simp] +lemma iterate_comp_eval : ∀ (k : ℕ) (t : R), + (p.comp^[k] q).eval t = ((λ x, p.eval x)^[k] (q.eval t)) := +iterate_comp_eval₂ _ + /-- `comp p`, regarded as a ring homomorphism from `R[X]` to itself. -/ def comp_ring_hom : R[X] → R[X] →+* R[X] := eval₂_ring_hom C From de87d5053a9fe5cbde723172c0fb7e27e7436473 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 27 Mar 2023 18:59:04 +0000 Subject: [PATCH 0705/1291] =?UTF-8?q?feat(algebra/order/monoid/min=5Fmax):?= =?UTF-8?q?=20`a=E2=82=81=20*=20b=E2=82=81=20=E2=89=A4=20a=E2=82=82=20*=20?= =?UTF-8?q?b=E2=82=82=20=E2=86=92=20a=E2=82=81=20=E2=89=A4=20a=E2=82=82=20?= =?UTF-8?q?=E2=88=A8=20b=E2=82=81=20=E2=89=A4=20b=E2=82=82`=20(#18667)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Complete the API. --- src/algebra/order/monoid/canonical/defs.lean | 2 + src/algebra/order/monoid/min_max.lean | 53 ++++++++++++-------- src/order/basic.lean | 3 ++ 3 files changed, 38 insertions(+), 20 deletions(-) diff --git a/src/algebra/order/monoid/canonical/defs.lean b/src/algebra/order/monoid/canonical/defs.lean index ffe6d64547c33..03ca16ab30062 100644 --- a/src/algebra/order/monoid/canonical/defs.lean +++ b/src/algebra/order/monoid/canonical/defs.lean @@ -123,6 +123,8 @@ variables [canonically_ordered_monoid α] {a b c d : α} @[to_additive] lemma le_of_mul_le_left : a * b ≤ c → a ≤ c := le_self_mul.trans @[to_additive] lemma le_of_mul_le_right : a * b ≤ c → b ≤ c := le_mul_self.trans +@[to_additive] lemma le_mul_of_le_left : a ≤ b → a ≤ b * c := le_self_mul.trans' +@[to_additive] lemma le_mul_of_le_right : a ≤ c → a ≤ b * c := le_mul_self.trans' @[to_additive] lemma le_iff_exists_mul : a ≤ b ↔ ∃ c, b = a * c := diff --git a/src/algebra/order/monoid/min_max.lean b/src/algebra/order/monoid/min_max.lean index eb97ace732b46..eebcdd00179a0 100644 --- a/src/algebra/order/monoid/min_max.lean +++ b/src/algebra/order/monoid/min_max.lean @@ -13,6 +13,8 @@ import algebra.order.monoid.lemmas > Any changes to this file require a corresponding PR to mathlib4. -/ +open function + variables {α β : Type*} /-! Some lemmas about types that have an ordering and a binary operation, with no @@ -43,26 +45,6 @@ variable [covariant_class α α (*) (≤)] lemma max_mul_mul_left (a b c : α) : max (a * b) (a * c) = a * max b c := (monotone_id.const_mul' a).map_max.symm -@[to_additive] -lemma lt_or_lt_of_mul_lt_mul [covariant_class α α (function.swap (*)) (≤)] - {a b m n : α} (h : m * n < a * b) : - m < a ∨ n < b := -by { contrapose! h, exact mul_le_mul' h.1 h.2 } - -@[to_additive] -lemma mul_lt_mul_iff_of_le_of_le - [covariant_class α α (function.swap (*)) (<)] - [covariant_class α α (*) (<)] - [covariant_class α α (function.swap (*)) (≤)] - {a b c d : α} (ac : a ≤ c) (bd : b ≤ d) : - a * b < c * d ↔ (a < c) ∨ (b < d) := -begin - refine ⟨lt_or_lt_of_mul_lt_mul, λ h, _⟩, - cases h with ha hb, - { exact mul_lt_mul_of_lt_of_le ha bd }, - { exact mul_lt_mul_of_le_of_lt ac hb } -end - end left section right @@ -78,6 +60,37 @@ lemma max_mul_mul_right (a b c : α) : max (a * c) (b * c) = max a b * c := end right +@[to_additive] lemma lt_or_lt_of_mul_lt_mul [covariant_class α α (*) (≤)] + [covariant_class α α (swap (*)) (≤)] {a₁ a₂ b₁ b₂ : α} : + a₁ * b₁ < a₂ * b₂ → a₁ < a₂ ∨ b₁ < b₂ := +by { contrapose!, exact λ h, mul_le_mul' h.1 h.2 } + +@[to_additive] lemma le_or_lt_of_mul_le_mul [covariant_class α α (*) (≤)] + [covariant_class α α (swap (*)) (<)] {a₁ a₂ b₁ b₂ : α} : + a₁ * b₁ ≤ a₂ * b₂ → a₁ ≤ a₂ ∨ b₁ < b₂ := +by { contrapose!, exact λ h, mul_lt_mul_of_lt_of_le h.1 h.2 } + +@[to_additive] lemma lt_or_le_of_mul_le_mul [covariant_class α α (*) (<)] + [covariant_class α α (swap (*)) (≤)] {a₁ a₂ b₁ b₂ : α} : + a₁ * b₁ ≤ a₂ * b₂ → a₁ < a₂ ∨ b₁ ≤ b₂ := +by { contrapose!, exact λ h, mul_lt_mul_of_le_of_lt h.1 h.2 } + +@[to_additive] lemma le_or_le_of_mul_le_mul [covariant_class α α (*) (<)] + [covariant_class α α (swap (*)) (<)] {a₁ a₂ b₁ b₂ : α} : + a₁ * b₁ ≤ a₂ * b₂ → a₁ ≤ a₂ ∨ b₁ ≤ b₂ := +by { contrapose!, exact λ h, mul_lt_mul_of_lt_of_lt h.1 h.2 } + +@[to_additive] lemma mul_lt_mul_iff_of_le_of_le [covariant_class α α (*) (≤)] + [covariant_class α α (swap (*)) (≤)] [covariant_class α α (*) (<)] + [covariant_class α α (swap (*)) (<)] {a₁ a₂ b₁ b₂ : α} (ha : a₁ ≤ a₂) (hb : b₁ ≤ b₂) : + a₁ * b₁ < a₂ * b₂ ↔ a₁ < a₂ ∨ b₁ < b₂ := +begin + refine ⟨lt_or_lt_of_mul_lt_mul, _⟩, + rintro (ha | hb), + { exact mul_lt_mul_of_lt_of_le ha hb }, + { exact mul_lt_mul_of_le_of_lt ha hb } +end + end has_mul variable [mul_one_class α] diff --git a/src/order/basic.lean b/src/order/basic.lean index 61feee4706316..708b8e3bdf4e2 100644 --- a/src/order/basic.lean +++ b/src/order/basic.lean @@ -239,10 +239,13 @@ by haveI := classical.dec; exact decidable.eq_iff_le_not_lt lemma eq_or_lt_of_le [partial_order α] {a b : α} (h : a ≤ b) : a = b ∨ a < b := h.lt_or_eq.symm lemma eq_or_gt_of_le [partial_order α] {a b : α} (h : a ≤ b) : b = a ∨ a < b := h.lt_or_eq.symm.imp eq.symm id +lemma gt_or_eq_of_le [partial_order α] {a b : α} (hab : a ≤ b) : a < b ∨ b = a := +(eq_or_gt_of_le hab).symm alias decidable.eq_or_lt_of_le ← has_le.le.eq_or_lt_dec alias eq_or_lt_of_le ← has_le.le.eq_or_lt alias eq_or_gt_of_le ← has_le.le.eq_or_gt +alias gt_or_eq_of_le ← has_le.le.gt_or_eq attribute [nolint decidable_classical] has_le.le.eq_or_lt_dec From b31173ee05c911d61ad6a05bd2196835c932e0ec Mon Sep 17 00:00:00 2001 From: Moritz Doll Date: Tue, 28 Mar 2023 01:31:58 +0000 Subject: [PATCH 0706/1291] chore(analysis/locally_convex/with_seminorms): remove unnecessary positivity assumption (#18659) In the boundedness definition it is not necessary to assume that the bound is positive. --- .../locally_convex/with_seminorms.lean | 40 +++++++++---------- src/analysis/schwartz_space.lean | 4 +- 2 files changed, 21 insertions(+), 23 deletions(-) diff --git a/src/analysis/locally_convex/with_seminorms.lean b/src/analysis/locally_convex/with_seminorms.lean index a39d6f380a131..a578827117d1d 100644 --- a/src/analysis/locally_convex/with_seminorms.lean +++ b/src/analysis/locally_convex/with_seminorms.lean @@ -221,46 +221,42 @@ variables {σ₁₂ : 𝕜 →+* 𝕜₂} [ring_hom_isometric σ₁₂] /-- The proposition that a linear map is bounded between spaces with families of seminorms. -/ def is_bounded (p : ι → seminorm 𝕜 E) (q : ι' → seminorm 𝕜₂ F) (f : E →ₛₗ[σ₁₂] F) : Prop := - ∀ i, ∃ s : finset ι, ∃ C : ℝ≥0, C ≠ 0 ∧ (q i).comp f ≤ C • s.sup p + ∀ i, ∃ s : finset ι, ∃ C : ℝ≥0, (q i).comp f ≤ C • s.sup p lemma is_bounded_const (ι' : Type*) [nonempty ι'] {p : ι → seminorm 𝕜 E} {q : seminorm 𝕜₂ F} (f : E →ₛₗ[σ₁₂] F) : - is_bounded p (λ _ : ι', q) f ↔ ∃ (s : finset ι) C : ℝ≥0, C ≠ 0 ∧ q.comp f ≤ C • s.sup p := + is_bounded p (λ _ : ι', q) f ↔ ∃ (s : finset ι) C : ℝ≥0, q.comp f ≤ C • s.sup p := by simp only [is_bounded, forall_const] lemma const_is_bounded (ι : Type*) [nonempty ι] {p : seminorm 𝕜 E} {q : ι' → seminorm 𝕜₂ F} (f : E →ₛₗ[σ₁₂] F) : - is_bounded (λ _ : ι, p) q f ↔ ∀ i, ∃ C : ℝ≥0, C ≠ 0 ∧ (q i).comp f ≤ C • p := + is_bounded (λ _ : ι, p) q f ↔ ∀ i, ∃ C : ℝ≥0, (q i).comp f ≤ C • p := begin split; intros h i, - { rcases h i with ⟨s, C, hC, h⟩, - exact ⟨C, hC, le_trans h (smul_le_smul (finset.sup_le (λ _ _, le_rfl)) le_rfl)⟩ }, + { rcases h i with ⟨s, C, h⟩, + exact ⟨C, le_trans h (smul_le_smul (finset.sup_le (λ _ _, le_rfl)) le_rfl)⟩ }, use [{classical.arbitrary ι}], simp only [h, finset.sup_singleton], end lemma is_bounded_sup {p : ι → seminorm 𝕜 E} {q : ι' → seminorm 𝕜₂ F} {f : E →ₛₗ[σ₁₂] F} (hf : is_bounded p q f) (s' : finset ι') : - ∃ (C : ℝ≥0) (s : finset ι), 0 < C ∧ (s'.sup q).comp f ≤ C • (s.sup p) := + ∃ (C : ℝ≥0) (s : finset ι), (s'.sup q).comp f ≤ C • (s.sup p) := begin classical, obtain rfl | hs' := s'.eq_empty_or_nonempty, - { exact ⟨1, ∅, zero_lt_one, by simp [seminorm.bot_eq_zero]⟩ }, + { exact ⟨1, ∅, by simp [seminorm.bot_eq_zero]⟩ }, choose fₛ fC hf using hf, use [s'.card • s'.sup fC, finset.bUnion s' fₛ], - split, - { refine nsmul_pos _ (ne_of_gt (finset.nonempty.card_pos hs')), - cases finset.nonempty.bex hs' with j hj, - exact lt_of_lt_of_le (zero_lt_iff.mpr (and.elim_left (hf j))) (finset.le_sup hj) }, have hs : ∀ i : ι', i ∈ s' → (q i).comp f ≤ s'.sup fC • ((finset.bUnion s' fₛ).sup p) := begin intros i hi, - refine le_trans (and.elim_right (hf i)) (smul_le_smul _ (finset.le_sup hi)), + refine (hf i).trans (smul_le_smul _ (finset.le_sup hi)), exact finset.sup_mono (finset.subset_bUnion_of_mem fₛ hi), end, - refine le_trans (comp_mono f (finset_sup_le_sum q s')) _, + refine (comp_mono f (finset_sup_le_sum q s')).trans _, simp_rw [←pullback_apply, add_monoid_hom.map_sum, pullback_apply], - refine le_trans (finset.sum_le_sum hs) _, + refine (finset.sum_le_sum hs).trans _, rw [finset.sum_const, smul_assoc], exact le_rfl, end @@ -608,20 +604,22 @@ begin refine continuous_of_continuous_comp hq _ (λ i, seminorm.continuous_of_continuous_at_zero _), rw [metric.continuous_at_iff', map_zero], intros r hr, - rcases hf i with ⟨s₁, C, hC, hf⟩, - have hC' : 0 < C := hC.bot_lt, + rcases hf i with ⟨s₁, C, hf⟩, + have hC' : 0 < C + 1 := by positivity, rw hp.has_basis.eventually_iff, - refine ⟨(s₁.sup p).ball 0 (r/C), p.basis_sets_mem _ (by positivity), _⟩, + refine ⟨(s₁.sup p).ball 0 (r/(C + 1)), p.basis_sets_mem _ (by positivity), _⟩, simp_rw [ ←metric.mem_ball, ←mem_preimage, ←ball_zero_eq_preimage_ball], refine subset.trans _ (ball_antitone hf), - rw ball_smul (s₁.sup p) hC', - refl + norm_cast, + rw ← ball_smul (s₁.sup p) hC', + refine ball_antitone (smul_le_smul le_rfl _), + simp only [le_add_iff_nonneg_right, zero_le'], end lemma cont_with_seminorms_normed_space (F) [seminormed_add_comm_group F] [normed_space 𝕝₂ F] [uniform_space E] [uniform_add_group E] {p : ι → seminorm 𝕝 E} (hp : with_seminorms p) (f : E →ₛₗ[τ₁₂] F) - (hf : ∃ (s : finset ι) C : ℝ≥0, C ≠ 0 ∧ (norm_seminorm 𝕝₂ F).comp f ≤ C • s.sup p) : + (hf : ∃ (s : finset ι) C : ℝ≥0, (norm_seminorm 𝕝₂ F).comp f ≤ C • s.sup p) : continuous f := begin rw ←seminorm.is_bounded_const (fin 1) at hf, @@ -631,7 +629,7 @@ end lemma cont_normed_space_to_with_seminorms (E) [seminormed_add_comm_group E] [normed_space 𝕝 E] [uniform_space F] [uniform_add_group F] {q : ι → seminorm 𝕝₂ F} (hq : with_seminorms q) (f : E →ₛₗ[τ₁₂] F) - (hf : ∀ i : ι, ∃ C : ℝ≥0, C ≠ 0 ∧ (q i).comp f ≤ C • (norm_seminorm 𝕝 E)) : continuous f := + (hf : ∀ i : ι, ∃ C : ℝ≥0, (q i).comp f ≤ C • (norm_seminorm 𝕝 E)) : continuous f := begin rw ←seminorm.const_is_bounded (fin 1) at hf, exact continuous_from_bounded (norm_with_seminorms 𝕝 E) hq f hf, diff --git a/src/analysis/schwartz_space.lean b/src/analysis/schwartz_space.lean index a29de8178f72d..0a2d35e132e9a 100644 --- a/src/analysis/schwartz_space.lean +++ b/src/analysis/schwartz_space.lean @@ -504,7 +504,7 @@ def fderiv_clm : 𝓢(E, F) →L[𝕜] 𝓢(E, E →L[ℝ] F) := refine seminorm.continuous_from_bounded (schwartz_with_seminorms 𝕜 E F) (schwartz_with_seminorms 𝕜 E (E →L[ℝ] F)) _ _, rintros ⟨k, n⟩, - use [{⟨k, n+1⟩}, 1, one_ne_zero], + use [{⟨k, n+1⟩}, 1], intros f, simp only [schwartz_seminorm_family_apply, seminorm.comp_apply, finset.sup_singleton, one_smul], refine (fderiv_lm 𝕜 f).seminorm_le_bound 𝕜 k n (by positivity) _, @@ -556,7 +556,7 @@ def to_bounded_continuous_function_clm : 𝓢(E, F) →L[𝕜] E →ᵇ F := begin change continuous (to_bounded_continuous_function_lm 𝕜 E F), refine seminorm.continuous_from_bounded (schwartz_with_seminorms 𝕜 E F) - (norm_with_seminorms 𝕜 (E →ᵇ F)) _ (λ i, ⟨{0}, 1, one_ne_zero, λ f, _⟩), + (norm_with_seminorms 𝕜 (E →ᵇ F)) _ (λ i, ⟨{0}, 1, λ f, _⟩), rw [finset.sup_singleton, one_smul , seminorm.comp_apply, coe_norm_seminorm, schwartz_seminorm_family_apply_zero, bounded_continuous_function.norm_le (map_nonneg _ _)], intros x, From ef5f2ce93dd79b7fb184018b6f48132a10c764e7 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Tue, 28 Mar 2023 07:28:37 +0000 Subject: [PATCH 0707/1291] feat(set_theory/zfc/basic): simpler/more general `Class.ext` (#18306) I failed to notice `set.ext` was simpler and more general than my custom `Class.ext` lemma. As such, I've replaced it. --- src/set_theory/zfc/basic.lean | 35 ++++++++++++++--------------------- 1 file changed, 14 insertions(+), 21 deletions(-) diff --git a/src/set_theory/zfc/basic.lean b/src/set_theory/zfc/basic.lean index e0e162b352924..ff429736a9045 100644 --- a/src/set_theory/zfc/basic.lean +++ b/src/set_theory/zfc/basic.lean @@ -921,6 +921,10 @@ def Class := set Set namespace Class +@[ext] theorem ext {x y : Class.{u}} : (∀ z : Set.{u}, x z ↔ y z) → x = y := set.ext + +theorem ext_iff {x y : Class.{u}} : x = y ↔ ∀ z, x z ↔ y z := set.ext_iff + /-- Coerce a ZFC set into a class -/ def of_Set (x : Set.{u}) : Class.{u} := {y | y ∈ x} instance : has_coe Set Class := ⟨of_Set⟩ @@ -1007,26 +1011,26 @@ to_Set_of_Set _ _ @[simp, norm_cast] theorem coe_sep (p : Class.{u}) (x : Set.{u}) : (↑{y ∈ x | p y} : Class.{u}) = {y ∈ x | p y} := -set.ext $ λ y, Set.mem_sep +ext $ λ y, Set.mem_sep @[simp, norm_cast] theorem coe_empty : ↑(∅ : Set.{u}) = (∅ : Class.{u}) := -set.ext $ λ y, (iff_false _).2 $ Set.not_mem_empty y +ext $ λ y, (iff_false _).2 $ Set.not_mem_empty y @[simp, norm_cast] theorem coe_insert (x y : Set.{u}) : ↑(insert x y) = @insert Set.{u} Class.{u} _ x y := -set.ext $ λ z, Set.mem_insert_iff +ext $ λ z, Set.mem_insert_iff @[simp, norm_cast] theorem coe_union (x y : Set.{u}) : ↑(x ∪ y) = (x : Class.{u}) ∪ y := -set.ext $ λ z, Set.mem_union +ext $ λ z, Set.mem_union @[simp, norm_cast] theorem coe_inter (x y : Set.{u}) : ↑(x ∩ y) = (x : Class.{u}) ∩ y := -set.ext $ λ z, Set.mem_inter +ext $ λ z, Set.mem_inter @[simp, norm_cast] theorem coe_diff (x y : Set.{u}) : ↑(x \ y) = (x : Class.{u}) \ y := -set.ext $ λ z, Set.mem_diff +ext $ λ z, Set.mem_diff @[simp, norm_cast] theorem coe_powerset (x : Set.{u}) : ↑x.powerset = powerset.{u} x := -set.ext $ λ z, Set.mem_powerset +ext $ λ z, Set.mem_powerset @[simp] theorem powerset_apply {A : Class.{u}} {x : Set.{u}} : powerset A x ↔ ↑x ⊆ A := iff.rfl @@ -1039,18 +1043,7 @@ begin end @[simp, norm_cast] theorem coe_sUnion (x : Set.{u}) : ↑(⋃₀ x) = ⋃₀ (x : Class.{u}) := -set.ext $ λ y, Set.mem_sUnion.trans (sUnion_apply.trans $ by simp_rw [coe_apply, exists_prop]).symm - -@[ext] theorem ext {x y : Class.{u}} : (∀ z : Class.{u}, z ∈ x ↔ z ∈ y) → x = y := -begin - refine λ h, set.ext (λ z, _), - change x z ↔ y z, - rw [←@coe_mem z x, ←@coe_mem z y], - exact h z -end - -theorem ext_iff {x y : Class.{u}} : x = y ↔ (∀ z : Class.{u}, z ∈ x ↔ z ∈ y) := -⟨λ h, by simp [h], ext⟩ +ext $ λ y, Set.mem_sUnion.trans (sUnion_apply.trans $ by simp_rw [coe_apply, exists_prop]).symm @[simp] theorem mem_sUnion {x y : Class.{u}} : y ∈ ⋃₀ x ↔ ∃ z, z ∈ x ∧ y ∈ z := begin @@ -1077,7 +1070,7 @@ end def iota (A : Class) : Class := ⋃₀ {x | ∀ y, A y ↔ y = x} theorem iota_val (A : Class) (x : Set) (H : ∀ y, A y ↔ y = x) : iota A = ↑x := -set.ext $ λ y, ⟨λ ⟨._, ⟨x', rfl, h⟩, yx'⟩, by rwa ←((H x').1 $ (h x').2 rfl), +ext $ λ y, ⟨λ ⟨._, ⟨x', rfl, h⟩, yx'⟩, by rwa ←((H x').1 $ (h x').2 rfl), λ yx, ⟨_, ⟨x, rfl, H⟩, yx⟩⟩ /-- Unlike the other set constructors, the `iota` definite descriptor @@ -1086,7 +1079,7 @@ set.ext $ λ y, ⟨λ ⟨._, ⟨x', rfl, h⟩, yx'⟩, by rwa ←((H x').1 $ (h theorem iota_ex (A) : iota.{u} A ∈ univ.{u} := mem_univ.2 $ or.elim (classical.em $ ∃ x, ∀ y, A y ↔ y = x) (λ ⟨x, h⟩, ⟨x, eq.symm $ iota_val A x h⟩) - (λ hn, ⟨∅, set.ext (λ z, coe_empty.symm ▸ ⟨false.rec _, λ ⟨._, ⟨x, rfl, H⟩, zA⟩, hn ⟨x, H⟩⟩)⟩) + (λ hn, ⟨∅, ext (λ z, coe_empty.symm ▸ ⟨false.rec _, λ ⟨._, ⟨x, rfl, H⟩, zA⟩, hn ⟨x, H⟩⟩)⟩) /-- Function value -/ def fval (F A : Class.{u}) : Class.{u} := iota (λ y, to_Set (λ x, F (Set.pair x y)) A) From 0148d455199ed64bf8eb2f493a1e7eb9211ce170 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Tue, 28 Mar 2023 11:22:01 +0000 Subject: [PATCH 0708/1291] chore(category_theory/adjunction/opposites): backport removal of simp lemmas (#18680) Some of the lemmas generated by `simps` in https://github.com/leanprover-community/mathlib4/pull/2424 are bad according to the simpNF linter, and have proved hard to fix by hand. Fortunately, they are simply not needed. This PR verifies this by backporting their removal to mathlib3. Compiles locally, lets hope CI agrees. Co-authored-by: Scott Morrison --- src/category_theory/adjunction/opposites.lean | 6 ++++-- 1 file changed, 4 insertions(+), 2 deletions(-) diff --git a/src/category_theory/adjunction/opposites.lean b/src/category_theory/adjunction/opposites.lean index e5c6476be72a8..8566a4b81d802 100644 --- a/src/category_theory/adjunction/opposites.lean +++ b/src/category_theory/adjunction/opposites.lean @@ -29,7 +29,8 @@ variables {C : Type u₁} [category.{v₁} C] {D : Type u₂} [category.{v₂} D namespace category_theory.adjunction /-- If `G.op` is adjoint to `F.op` then `F` is adjoint to `G`. -/ -@[simps] def adjoint_of_op_adjoint_op (F : C ⥤ D) (G : D ⥤ C) (h : G.op ⊣ F.op) : F ⊣ G := +@[simps unit_app counit_app] def adjoint_of_op_adjoint_op + (F : C ⥤ D) (G : D ⥤ C) (h : G.op ⊣ F.op) : F ⊣ G := adjunction.mk_of_hom_equiv { hom_equiv := λ X Y, ((h.hom_equiv (opposite.op Y) (opposite.op X)).trans (op_equiv _ _)).symm.trans (op_equiv _ _) } @@ -47,7 +48,8 @@ def unop_adjoint_unop_of_adjoint (F : Cᵒᵖ ⥤ Dᵒᵖ) (G : Dᵒᵖ ⥤ Cᵒ adjoint_unop_of_adjoint_op F.unop G (h.of_nat_iso_right F.op_unop_iso.symm) /-- If `G` is adjoint to `F` then `F.op` is adjoint to `G.op`. -/ -@[simps] def op_adjoint_op_of_adjoint (F : C ⥤ D) (G : D ⥤ C) (h : G ⊣ F) : F.op ⊣ G.op := +@[simps unit_app counit_app] def op_adjoint_op_of_adjoint + (F : C ⥤ D) (G : D ⥤ C) (h : G ⊣ F) : F.op ⊣ G.op := adjunction.mk_of_hom_equiv { hom_equiv := λ X Y, (op_equiv _ Y).trans ((h.hom_equiv _ _).symm.trans (op_equiv X (opposite.op _)).symm) } From 160f568dcf772b2477791c844fc605f2f91f73d1 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Tue, 28 Mar 2023 21:02:25 +0000 Subject: [PATCH 0709/1291] feat(algebraic_topology/dold_kan): tools for compatibilities (#17922) Tools are introduced in order to construct Dold-Kan equivalences with good definitional properties. --- .../dold_kan/compatibility.lean | 183 ++++++++++++++++++ 1 file changed, 183 insertions(+) create mode 100644 src/algebraic_topology/dold_kan/compatibility.lean diff --git a/src/algebraic_topology/dold_kan/compatibility.lean b/src/algebraic_topology/dold_kan/compatibility.lean new file mode 100644 index 0000000000000..8093d7bfa79ce --- /dev/null +++ b/src/algebraic_topology/dold_kan/compatibility.lean @@ -0,0 +1,183 @@ +/- +Copyright (c) 2022 Joël Riou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joël Riou +-/ + +import category_theory.equivalence + +/-! Tools for compatibilities between Dold-Kan equivalences + +The purpose of this file is to introduce tools which will enable the +construction of the Dold-Kan equivalence `simplicial_object C ≌ chain_complex C ℕ` +for a pseudoabelian category `C` from the equivalence +`karoubi (simplicial_object C) ≌ karoubi (chain_complex C ℕ)` and the two +equivalences `simplicial_object C ≅ karoubi (simplicial_object C)` and +`chain_complex C ℕ ≅ karoubi (chain_complex C ℕ)`. + +It is certainly possible to get an equivalence `simplicial_object C ≌ chain_complex C ℕ` +using a compositions of the three equivalences above, but then neither the functor +nor the inverse would have good definitional properties. For example, it would be better +if the inverse functor of the equivalence was exactly the functor +`Γ₀ : simplicial_object C ⥤ chain_complex C ℕ` which was constructed in `functor_gamma.lean`. + +In this file, given four categories `A`, `A'`, `B`, `B'`, equivalences `eA : A ≅ A'`, +`eB : B ≅ B'`, `e' : A' ≅ B'`, functors `F : A ⥤ B'`, `G : B ⥤ A` equipped with certain +compatibilities, we construct successive equivalences: +- `equivalence₀` from `A` to `B'`, which is the composition of `eA` and `e'`. +- `equivalence₁` from `A` to `B'`, with the same inverse functor as `equivalence₀`, +but whose functor is `F`. +- `equivalence₂` from `A` to `B`, which is the composition of `equivalence₁` and the +inverse of `eB`: +- `equivalence` from `A` to `B`, which has the same functor `F ⋙ eB.inverse` as `equivalence₂`, +but whose inverse functor is `G`. + +When extra assumptions are given, we shall also provide simplification lemmas for the +unit and counit isomorphisms of `equivalence`. (TODO) + +-/ + +open category_theory category_theory.category + +namespace algebraic_topology + +namespace dold_kan + +namespace compatibility + +variables {A A' B B' : Type*} [category A] [category A'] [category B] [category B'] + (eA : A ≌ A') (eB : B ≌ B') (e' : A' ≌ B') + {F : A ⥤ B'} (hF : eA.functor ⋙ e'.functor ≅ F) + {G : B ⥤ A} (hG : eB.functor ⋙ e'.inverse ≅ G ⋙ eA.functor) + +/-- A basic equivalence `A ≅ B'` obtained by composing `eA : A ≅ A'` and `e' : A' ≅ B'`. -/ +@[simps functor inverse unit_iso_hom_app] +def equivalence₀ : A ≌ B' := eA.trans e' + +include hF +variables {eA} {e'} + +/-- An intermediate equivalence `A ≅ B'` whose functor is `F` and whose inverse is +`e'.inverse ⋙ eA.inverse`. -/ +@[simps functor] +def equivalence₁ : A ≌ B' := +begin + letI : is_equivalence F := + is_equivalence.of_iso hF (is_equivalence.of_equivalence (equivalence₀ eA e')), + exact F.as_equivalence, +end + +lemma equivalence₁_inverse : (equivalence₁ hF).inverse = e'.inverse ⋙ eA.inverse := rfl + +/-- The counit isomorphism of the equivalence `equivalence₁` between `A` and `B'`. -/ +@[simps] +def equivalence₁_counit_iso : + (e'.inverse ⋙ eA.inverse) ⋙ F ≅ 𝟭 B' := +calc (e'.inverse ⋙ eA.inverse) ⋙ F + ≅ (e'.inverse ⋙ eA.inverse) ⋙ (eA.functor ⋙ e'.functor) : iso_whisker_left _ hF.symm +... ≅ e'.inverse ⋙ (eA.inverse ⋙ eA.functor) ⋙ e'.functor : iso.refl _ +... ≅ e'.inverse ⋙ 𝟭 _ ⋙ e'.functor : iso_whisker_left _ (iso_whisker_right eA.counit_iso _) +... ≅ e'.inverse ⋙ e'.functor : iso.refl _ +... ≅ 𝟭 B' : e'.counit_iso + +lemma equivalence₁_counit_iso_eq : (equivalence₁ hF).counit_iso = equivalence₁_counit_iso hF := +begin + ext Y, + dsimp [equivalence₀, equivalence₁, is_equivalence.inverse, is_equivalence.of_equivalence], + simp only [equivalence₁_counit_iso_hom_app, category_theory.functor.map_id, comp_id], +end + +/-- The unit isomorphism of the equivalence `equivalence₁` between `A` and `B'`. -/ +@[simps] +def equivalence₁_unit_iso : + 𝟭 A ≅ F ⋙ (e'.inverse ⋙ eA.inverse) := +calc 𝟭 A ≅ eA.functor ⋙ eA.inverse : eA.unit_iso +... ≅ eA.functor ⋙ 𝟭 A' ⋙ eA.inverse : iso.refl _ +... ≅ eA.functor ⋙ (e'.functor ⋙ e'.inverse) ⋙ eA.inverse : + iso_whisker_left _ (iso_whisker_right e'.unit_iso _) +... ≅ (eA.functor ⋙ e'.functor) ⋙ (e'.inverse ⋙ eA.inverse) : iso.refl _ +... ≅ F ⋙ (e'.inverse ⋙ eA.inverse) : iso_whisker_right hF _ + +lemma equivalence₁_unit_iso_eq : (equivalence₁ hF).unit_iso = equivalence₁_unit_iso hF := +begin + ext X, + dsimp [equivalence₀, equivalence₁, nat_iso.hcomp, + is_equivalence.of_equivalence], + simp only [id_comp, assoc, equivalence₁_unit_iso_hom_app], +end + +include eB + +/-- An intermediate equivalence `A ≅ B` obtained as the composition of `equivalence₁` and +the inverse of `eB : B ≌ B'`. -/ +@[simps functor] +def equivalence₂ : A ≌ B := (equivalence₁ hF).trans eB.symm + +lemma equivalence₂_inverse : (equivalence₂ eB hF).inverse = + eB.functor ⋙ e'.inverse ⋙ eA.inverse := rfl + +/-- The counit isomorphism of the equivalence `equivalence₂` between `A` and `B`. -/ +@[simps] +def equivalence₂_counit_iso : + (eB.functor ⋙ e'.inverse ⋙ eA.inverse) ⋙ (F ⋙ eB.inverse) ≅ 𝟭 B := +calc (eB.functor ⋙ e'.inverse ⋙ eA.inverse) ⋙ (F ⋙ eB.inverse) + ≅ eB.functor ⋙ (e'.inverse ⋙ eA.inverse ⋙ F) ⋙ eB.inverse : iso.refl _ +... ≅ eB.functor ⋙ 𝟭 _ ⋙ eB.inverse : + iso_whisker_left _ (iso_whisker_right (equivalence₁_counit_iso hF) _) +... ≅ eB.functor ⋙ eB.inverse : iso.refl _ +... ≅ 𝟭 B : eB.unit_iso.symm + +lemma equivalence₂_counit_iso_eq : + (equivalence₂ eB hF).counit_iso = equivalence₂_counit_iso eB hF := +begin + ext Y', + dsimp [equivalence₂, iso.refl], + simp only [equivalence₁_counit_iso_eq, equivalence₂_counit_iso_hom_app, + equivalence₁_counit_iso_hom_app, functor.map_comp, assoc], +end + +/-- The unit isomorphism of the equivalence `equivalence₂` between `A` and `B`. -/ +@[simps] +def equivalence₂_unit_iso : + 𝟭 A ≅ (F ⋙ eB.inverse) ⋙ (eB.functor ⋙ e'.inverse ⋙ eA.inverse) := +calc 𝟭 A ≅ F ⋙ e'.inverse ⋙ eA.inverse : equivalence₁_unit_iso hF +... ≅ F ⋙ 𝟭 B' ⋙ (e'.inverse ⋙ eA.inverse) : iso.refl _ +... ≅ F ⋙ (eB.inverse ⋙ eB.functor) ⋙ e'.inverse ⋙ eA.inverse : + iso_whisker_left _ (iso_whisker_right eB.counit_iso.symm _) +... ≅ (F ⋙ eB.inverse) ⋙ (eB.functor ⋙ e'.inverse ⋙ eA.inverse) : iso.refl _ + +lemma equivalence₂_unit_iso_eq : + (equivalence₂ eB hF).unit_iso = equivalence₂_unit_iso eB hF := +begin + ext X, + dsimp [equivalence₂], + simpa only [equivalence₂_unit_iso_hom_app, equivalence₁_unit_iso_eq, + equivalence₁_unit_iso_hom_app, assoc, nat_iso.cancel_nat_iso_hom_left], +end + +variable {eB} +include hG + +/-- The equivalence `A ≅ B` whose functor is `F ⋙ eB.inverse` and +whose inverse is `G : B ≅ A`. -/ +@[simps inverse] +def equivalence : A ≌ B := +begin + letI : is_equivalence G := begin + refine is_equivalence.of_iso _ (is_equivalence.of_equivalence (equivalence₂ eB hF).symm), + calc eB.functor ⋙ e'.inverse ⋙ eA.inverse + ≅ (eB.functor ⋙ e'.inverse) ⋙ eA.inverse : iso.refl _ + ... ≅ (G ⋙ eA.functor) ⋙ eA.inverse : iso_whisker_right hG _ + ... ≅ G ⋙ 𝟭 A : iso_whisker_left _ eA.unit_iso.symm + ... ≅ G : functor.right_unitor G, + end, + exact G.as_equivalence.symm, +end + +lemma equivalence_functor : (equivalence hF hG).functor = F ⋙ eB.inverse := rfl + +end compatibility + +end dold_kan + +end algebraic_topology From b99e2d58a5e6861833fa8de11e51a81144258db4 Mon Sep 17 00:00:00 2001 From: Anand Rao Date: Tue, 28 Mar 2023 22:29:20 +0000 Subject: [PATCH 0710/1291] feat(combinatorics/simple_graph/connected): support of connected components (#18442) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Rémi Bottinelli Co-authored-by: ART Co-authored-by: Anand Rao <18333981+0art0@users.noreply.github.com> --- .../simple_graph/connectivity.lean | 103 ++++++++++++++++++ src/combinatorics/simple_graph/ends/defs.lean | 2 +- 2 files changed, 104 insertions(+), 1 deletion(-) diff --git a/src/combinatorics/simple_graph/connectivity.lean b/src/combinatorics/simple_graph/connectivity.lean index 4ddd62eb75026..d6621ff6c91d3 100644 --- a/src/combinatorics/simple_graph/connectivity.lean +++ b/src/combinatorics/simple_graph/connectivity.lean @@ -1526,6 +1526,14 @@ protected lemma reachable.map {G : simple_graph V} {G' : simple_graph V'} (f : G →g G') {u v : V} (h : G.reachable u v) : G'.reachable (f u) (f v) := h.elim (λ p, ⟨p.map f⟩) +lemma iso.reachable_iff {G : simple_graph V} {G' : simple_graph V'} + {φ : G ≃g G'} {u v : V} : G'.reachable (φ u) (φ v) ↔ G.reachable u v := +⟨λ r, (φ.left_inv u) ▸ (φ.left_inv v) ▸ (r.map φ.symm.to_hom), reachable.map φ.to_hom⟩ + +lemma iso.symm_apply_reachable {G : simple_graph V} {G' : simple_graph V'} + {φ : G ≃g G'} {u : V} {v : V'} : G.reachable (φ.symm v) u ↔ G'.reachable v (φ u) := +by rw [← iso.reachable_iff, rel_iso.apply_symm_apply] + variables (G) lemma reachable_is_equivalence : equivalence G.reachable := @@ -1605,6 +1613,10 @@ protected lemma exact {v w : V} : G.connected_component_mk v = G.connected_component_mk w ↔ G.reachable v w := @quotient.eq _ G.reachable_setoid _ _ +lemma connected_component_mk_eq_of_adj {v w : V} (a : G.adj v w) : + G.connected_component_mk v = G.connected_component_mk w := +connected_component.sound a.reachable + /-- The `connected_component` specialization of `quot.lift`. Provides the stronger assumption that the vertices are connected by a path. -/ protected def lift {β : Sort*} (f : V → β) @@ -1642,6 +1654,96 @@ by { refine C.ind _, exact (λ _, rfl) } (φ : G →g G') (ψ : G' →g G'') : (C.map φ).map ψ = C.map (ψ.comp φ) := by { refine C.ind _, exact (λ _, rfl), } + +variables {φ : G ≃g G'} {v : V} {v' : V'} + +@[simp] lemma iso_image_comp_eq_map_iff_eq_comp + {C : G.connected_component} : + G'.connected_component_mk (φ v) = C.map (↑(↑φ : G ↪g G')) ↔ (G.connected_component_mk v) = C := +begin + refine C.ind (λ u, _), + simp only [iso.reachable_iff, connected_component.map_mk, + rel_embedding.coe_coe_fn, rel_iso.coe_coe_fn, connected_component.eq], +end + +@[simp] lemma iso_inv_image_comp_eq_iff_eq_map + {C : G.connected_component} : + G.connected_component_mk (φ.symm v') = C ↔ + G'.connected_component_mk v' = C.map φ := +begin + refine C.ind (λ u, _), + simp only [iso.symm_apply_reachable, connected_component.eq, coe_coe, + connected_component.map_mk, rel_embedding.coe_coe_fn, rel_iso.coe_coe_fn], +end + +end connected_component + +namespace iso + +/-- An isomorphism of graphs induces a bijection of connected components. -/ +@[simps] +def connected_component_equiv (φ : G ≃g G') : G.connected_component ≃ G'.connected_component := +{ to_fun := connected_component.map φ, + inv_fun := connected_component.map φ.symm, + left_inv := λ C, connected_component.ind + (λ v, congr_arg (G.connected_component_mk) (equiv.left_inv φ.to_equiv v)) C, + right_inv := λ C, connected_component.ind + (λ v, congr_arg (G'.connected_component_mk) (equiv.right_inv φ.to_equiv v)) C } + +@[simp] lemma connected_component_equiv_refl : + (iso.refl : G ≃g G).connected_component_equiv = equiv.refl _ := +by { ext ⟨v⟩, refl, } + +@[simp] lemma connected_component_equiv_symm (φ : G ≃g G') : + φ.symm.connected_component_equiv = φ.connected_component_equiv.symm := by { ext ⟨_⟩, refl, } + +@[simp] lemma connected_component_equiv_trans (φ : G ≃g G') (φ' : G' ≃g G'') : + connected_component_equiv (φ.trans φ') = + φ.connected_component_equiv.trans φ'.connected_component_equiv := by { ext ⟨_⟩, refl, } + +end iso + +namespace connected_component + +/-- The set of vertices in a connected component of a graph. -/ +def supp (C : G.connected_component) := + { v | G.connected_component_mk v = C } + +@[ext] lemma supp_injective : + function.injective (connected_component.supp : G.connected_component → set V) := +begin + refine connected_component.ind₂ _, + intros v w, + simp only [connected_component.supp, set.ext_iff, connected_component.eq, set.mem_set_of_eq], + intro h, + rw [reachable_comm, h], +end + +@[simp] +lemma supp_inj {C D : G.connected_component} : C.supp = D.supp ↔ C = D := +connected_component.supp_injective.eq_iff + +instance : set_like G.connected_component V := +{ coe := connected_component.supp, + coe_injective' := connected_component.supp_injective, } + +@[simp] lemma mem_supp_iff (C : G.connected_component) (v : V) : + v ∈ C.supp ↔ G.connected_component_mk v = C := iff.rfl + +lemma connected_component_mk_mem {v : V} : + v ∈ G.connected_component_mk v := by exact rfl + +/-- +The equivalence between connected components, induced by an isomorphism of graphs, +itself defines an equivalence on the supports of each connected component. +-/ +def iso_equiv_supp (φ : G ≃g G') (C : G.connected_component) : + C.supp ≃ (φ.connected_component_equiv C).supp := +{ to_fun := λ v, ⟨φ v, connected_component.iso_image_comp_eq_map_iff_eq_comp.mpr v.prop⟩, + inv_fun := λ v', ⟨φ.symm v', connected_component.iso_inv_image_comp_eq_iff_eq_map.mpr v'.prop⟩, + left_inv := λ v, subtype.ext_val (φ.to_equiv.left_inv ↑v), + right_inv := λ v, subtype.ext_val (φ.to_equiv.right_inv ↑v), } + end connected_component /-- A subgraph is connected if it is connected as a simple graph. -/ @@ -1994,3 +2096,4 @@ sym2.ind (λ v w, is_bridge_iff_adj_and_forall_cycle_not_mem) e end bridge_edges end simple_graph + diff --git a/src/combinatorics/simple_graph/ends/defs.lean b/src/combinatorics/simple_graph/ends/defs.lean index 6faf6eb580573..3f5c634f5f3d1 100644 --- a/src/combinatorics/simple_graph/ends/defs.lean +++ b/src/combinatorics/simple_graph/ends/defs.lean @@ -45,7 +45,7 @@ end lemma component_compl.supp_inj {C D : G.component_compl K} : C.supp = D.supp ↔ C = D := component_compl.supp_injective.eq_iff -instance : set_like (G.component_compl K) V := +instance component_compl.set_like : set_like (G.component_compl K) V := { coe := component_compl.supp, coe_injective' := λ C D, (component_compl.supp_inj).mp, } From 0187644979f2d3e10a06e916a869c994facd9a87 Mon Sep 17 00:00:00 2001 From: Floris van Doorn Date: Tue, 28 Mar 2023 23:45:14 +0000 Subject: [PATCH 0711/1291] feat(geometry/manifold/vector_bundle/tangent): use the new tangent bundle definition in the library (#18601) * Stop using the old tangent bundle definition, and use the new one instead * This affects `geometry.manifold.mfderiv` and later files. * We remove `cont_mdiff_at_iff_target`, which doesn't hold anymore in the new setting. * We prove `cont_mdiff_at_total_space`, which characterizes `C^n` maps into a vector bundle. We need a bunch of basic lemmas to prove this. This makes it easy to prove `smooth_zero_section`. * We move some results from `cont_mdiff_mfderiv` to `manifold/vector_bundle/basic` --- src/analysis/calculus/cont_diff.lean | 2 +- src/geometry/manifold/cont_mdiff.lean | 21 + src/geometry/manifold/cont_mdiff_mfderiv.lean | 243 ++----- src/geometry/manifold/mfderiv.lean | 88 ++- src/geometry/manifold/tangent_bundle.lean | 651 ------------------ .../manifold/vector_bundle/basic.lean | 147 +++- .../manifold/vector_bundle/tangent.lean | 67 +- src/topology/fiber_bundle/basic.lean | 45 +- src/topology/vector_bundle/basic.lean | 33 +- 9 files changed, 387 insertions(+), 910 deletions(-) delete mode 100644 src/geometry/manifold/tangent_bundle.lean diff --git a/src/analysis/calculus/cont_diff.lean b/src/analysis/calculus/cont_diff.lean index d4518001e7b6d..a0b45dce2eb49 100644 --- a/src/analysis/calculus/cont_diff.lean +++ b/src/analysis/calculus/cont_diff.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2019 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Sébastien Gouëzel +Authors: Sébastien Gouëzel, Floris van Doorn -/ import analysis.calculus.cont_diff_def import analysis.calculus.mean_value diff --git a/src/geometry/manifold/cont_mdiff.lean b/src/geometry/manifold/cont_mdiff.lean index db581e63c8753..c0e101cfb1241 100644 --- a/src/geometry/manifold/cont_mdiff.lean +++ b/src/geometry/manifold/cont_mdiff.lean @@ -1381,6 +1381,11 @@ begin { simp only with mfld_simps } end +lemma cont_mdiff_within_at.fst {f : N → M × M'} {s : set N} {x : N} + (hf : cont_mdiff_within_at J (I.prod I') n f s x) : + cont_mdiff_within_at J I n (λ x, (f x).1) s x := +cont_mdiff_within_at_fst.comp x hf (maps_to_image f s) + lemma cont_mdiff_at_fst {p : M × N} : cont_mdiff_at (I.prod J) I n prod.fst p := cont_mdiff_within_at_fst @@ -1436,6 +1441,11 @@ begin { simp only with mfld_simps } end +lemma cont_mdiff_within_at.snd {f : N → M × M'} {s : set N} {x : N} + (hf : cont_mdiff_within_at J (I.prod I') n f s x) : + cont_mdiff_within_at J I' n (λ x, (f x).2) s x := +cont_mdiff_within_at_snd.comp x hf (maps_to_image f s) + lemma cont_mdiff_at_snd {p : M × N} : cont_mdiff_at (I.prod J) J n prod.snd p := cont_mdiff_within_at_snd @@ -1494,6 +1504,17 @@ smooth_fst.fst.prod_mk $ smooth_fst.snd.prod_mk smooth_snd end projections +lemma cont_mdiff_within_at_prod_iff (f : M → M' × N') {s : set M} {x : M} : + cont_mdiff_within_at I (I'.prod J') n f s x ↔ + cont_mdiff_within_at I I' n (prod.fst ∘ f) s x ∧ + cont_mdiff_within_at I J' n (prod.snd ∘ f) s x := +by { refine ⟨λ h, ⟨h.fst, h.snd⟩, λ h, _⟩, simpa only [prod.mk.eta] using h.1.prod_mk h.2 } + +lemma cont_mdiff_at_prod_iff (f : M → M' × N') {x : M} : + cont_mdiff_at I (I'.prod J') n f x ↔ + cont_mdiff_at I I' n (prod.fst ∘ f) x ∧ cont_mdiff_at I J' n (prod.snd ∘ f) x := +by { simp_rw [← cont_mdiff_within_at_univ], exact cont_mdiff_within_at_prod_iff f } + section prod_map variables {g : N → N'} {r : set N} {y : N} diff --git a/src/geometry/manifold/cont_mdiff_mfderiv.lean b/src/geometry/manifold/cont_mdiff_mfderiv.lean index 7ef6055e3c23f..b002d3ee823d9 100644 --- a/src/geometry/manifold/cont_mdiff_mfderiv.lean +++ b/src/geometry/manifold/cont_mdiff_mfderiv.lean @@ -1,9 +1,8 @@ /- Copyright (c) 2020 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Sébastien Gouëzel +Authors: Sébastien Gouëzel, Floris van Doorn -/ -import geometry.manifold.cont_mdiff import geometry.manifold.mfderiv /-! @@ -20,8 +19,8 @@ and related notions. of a `Cⁿ` function is `Cᵐ` when `m + 1 ≤ n`. -/ -open set function filter charted_space smooth_manifold_with_corners -open_locale topology manifold +open set function filter charted_space smooth_manifold_with_corners bundle +open_locale topology manifold bundle /-! ### Definition of smooth functions between manifolds -/ @@ -42,6 +41,9 @@ variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] {F' : Type*} [normed_add_comm_group F'] [normed_space 𝕜 F'] {G' : Type*} [topological_space G'] {J' : model_with_corners 𝕜 F' G'} {N' : Type*} [topological_space N'] [charted_space G' N'] [J's : smooth_manifold_with_corners J' N'] +-- declare some additional normed spaces, used for fibers of vector bundles +{F₁ : Type*} [normed_add_comm_group F₁] [normed_space 𝕜 F₁] +{F₂ : Type*} [normed_add_comm_group F₂] [normed_space 𝕜 F₂] -- declare functions, sets, points and smoothness indices {f f₁ : M → M'} {s s₁ t : set M} {x : M} {m n : ℕ∞} @@ -104,7 +106,7 @@ space are model spaces in models with corners. The general fact is proved in lemma cont_mdiff_on.continuous_on_tangent_map_within_aux {f : H → H'} {s : set H} (hf : cont_mdiff_on I I' n f s) (hn : 1 ≤ n) (hs : unique_mdiff_on I s) : - continuous_on (tangent_map_within I I' f s) ((tangent_bundle.proj I H) ⁻¹' s) := + continuous_on (tangent_map_within I I' f s) (π (tangent_space I) ⁻¹' s) := begin suffices h : continuous_on (λ (p : H × E), (f p.fst, (fderiv_within 𝕜 (written_in_ext_chart_at I I' p.fst f) (I.symm ⁻¹' s ∩ range I) @@ -114,7 +116,7 @@ begin have B := ((tangent_bundle_model_space_homeomorph H' I').symm.continuous.comp_continuous_on h) .comp' A, have : (univ ∩ ⇑(tangent_bundle_model_space_homeomorph H I) ⁻¹' (prod.fst ⁻¹' s)) = - tangent_bundle.proj I H ⁻¹' s, + π (tangent_space I) ⁻¹' s, by { ext ⟨x, v⟩, simp only with mfld_simps }, rw this at B, apply B.congr, @@ -163,8 +165,7 @@ are model spaces in models with corners. The general fact is proved in lemma cont_mdiff_on.cont_mdiff_on_tangent_map_within_aux {f : H → H'} {s : set H} (hf : cont_mdiff_on I I' n f s) (hmn : m + 1 ≤ n) (hs : unique_mdiff_on I s) : - cont_mdiff_on I.tangent I'.tangent m (tangent_map_within I I' f s) - ((tangent_bundle.proj I H) ⁻¹' s) := + cont_mdiff_on I.tangent I'.tangent m (tangent_map_within I I' f s) (π (tangent_space I) ⁻¹' s) := begin have m_le_n : m ≤ n, { apply le_trans _ hmn, @@ -182,7 +183,7 @@ begin refine ⟨hf.continuous_on_tangent_map_within_aux one_le_n hs, λp q, _⟩, have A : range I ×ˢ univ ∩ ((equiv.sigma_equiv_prod H E).symm ∘ λ (p : E × E), ((I.symm) p.fst, p.snd)) ⁻¹' - (tangent_bundle.proj I H ⁻¹' s) + (π (tangent_space I) ⁻¹' s) = (range I ∩ I.symm ⁻¹' s) ×ˢ univ, by { ext ⟨x, v⟩, simp only with mfld_simps }, suffices h : cont_diff_on 𝕜 m (((λ (p : H' × E'), (I' p.fst, p.snd)) ∘ @@ -226,7 +227,7 @@ is `C^m` when `m+1 ≤ n`. -/ theorem cont_mdiff_on.cont_mdiff_on_tangent_map_within (hf : cont_mdiff_on I I' n f s) (hmn : m + 1 ≤ n) (hs : unique_mdiff_on I s) : cont_mdiff_on I.tangent I'.tangent m (tangent_map_within I I' f s) - ((tangent_bundle.proj I M) ⁻¹' s) := + (π (tangent_space I) ⁻¹' s) := begin /- The strategy of the proof is to avoid unfolding the definitions, and reduce by functoriality to the case of functions on the model spaces, where we have already proved the result. @@ -256,28 +257,28 @@ begin /- First step: local reduction on the space, to a set `s'` which is contained in chart domains. -/ refine cont_mdiff_on_of_locally_cont_mdiff_on (λp hp, _), have hf' := cont_mdiff_on_iff.1 hf, - simp [tangent_bundle.proj] at hp, - let l := chart_at H p.1, + simp only with mfld_simps at hp, + let l := chart_at H p.proj, set Dl := chart_at (model_prod H E) p with hDl, - let r := chart_at H' (f p.1), + let r := chart_at H' (f p.proj), let Dr := chart_at (model_prod H' E') (tangent_map_within I I' f s p), let il := chart_at (model_prod H E) (tangent_map I I l p), let ir := chart_at (model_prod H' E') (tangent_map I I' (r ∘ f) p), let s' := f ⁻¹' r.source ∩ s ∩ l.source, - let s'_lift := (tangent_bundle.proj I M)⁻¹' s', + let s'_lift := π (tangent_space I) ⁻¹' s', let s'l := l.target ∩ l.symm ⁻¹' s', - let s'l_lift := (tangent_bundle.proj I H) ⁻¹' s'l, + let s'l_lift := π (tangent_space I) ⁻¹' s'l, rcases continuous_on_iff'.1 hf'.1 r.source r.open_source with ⟨o, o_open, ho⟩, suffices h : cont_mdiff_on I.tangent I'.tangent m (tangent_map_within I I' f s) s'_lift, - { refine ⟨(tangent_bundle.proj I M)⁻¹' (o ∩ l.source), _, _, _⟩, - show is_open ((tangent_bundle.proj I M)⁻¹' (o ∩ l.source)), from - (is_open.inter o_open l.open_source).preimage (tangent_bundle_proj_continuous _ _) , - show p ∈ tangent_bundle.proj I M ⁻¹' (o ∩ l.source), - { simp [tangent_bundle.proj] at ⊢, - have : p.1 ∈ f ⁻¹' r.source ∩ s, by simp [hp], + { refine ⟨π (tangent_space I) ⁻¹' (o ∩ l.source), _, _, _⟩, + show is_open (π (tangent_space I) ⁻¹' (o ∩ l.source)), from + (is_open.inter o_open l.open_source).preimage (continuous_proj E _) , + show p ∈ π (tangent_space I) ⁻¹' (o ∩ l.source), + { simp, + have : p.proj ∈ f ⁻¹' r.source ∩ s, by simp [hp], rw ho at this, exact this.1 }, - { have : tangent_bundle.proj I M ⁻¹' s ∩ tangent_bundle.proj I M ⁻¹' (o ∩ l.source) = s'_lift, + { have : π (tangent_space I) ⁻¹' s ∩ π (tangent_space I) ⁻¹' (o ∩ l.source) = s'_lift, { dsimp only [s'_lift, s'], rw [ho], mfld_set_tac }, rw this, exact h } }, @@ -317,8 +318,8 @@ begin { have A : cont_mdiff_on I'.tangent I'.tangent m Dr.symm Dr.target := cont_mdiff_on_chart_symm, apply cont_mdiff_on.comp A diff_irrfl_lift (λp hp, _), - simp only [s'l_lift, tangent_bundle.proj] with mfld_simps at hp, - simp only [ir, @local_equiv.refl_coe (model_prod H' E'), hp] with mfld_simps }, + simp only [s'l_lift] with mfld_simps at hp, + simp only [ir, hp] with mfld_simps }, -- conclusion of this step: the composition of all the maps above is smooth have diff_DrirrflilDl : cont_mdiff_on I.tangent I'.tangent m (Dr.symm ∘ (ir ∘ (tangent_map_within I I' (r ∘ f ∘ l.symm) s'l)) ∘ @@ -326,22 +327,22 @@ begin { have A : cont_mdiff_on I.tangent I.tangent m Dl Dl.source := cont_mdiff_on_chart, have A' : cont_mdiff_on I.tangent I.tangent m Dl s'_lift, { apply A.mono (λp hp, _), - simp only [s'_lift, tangent_bundle.proj] with mfld_simps at hp, + simp only [s'_lift] with mfld_simps at hp, simp only [Dl, hp] with mfld_simps }, have B : cont_mdiff_on I.tangent I.tangent m il.symm il.target := cont_mdiff_on_chart_symm, have C : cont_mdiff_on I.tangent I.tangent m (il.symm ∘ Dl) s'_lift := cont_mdiff_on.comp B A' (λp hp, by simp only [il] with mfld_simps), apply cont_mdiff_on.comp diff_Drirrfl_lift C (λp hp, _), - simp only [s'_lift, tangent_bundle.proj] with mfld_simps at hp, - simp only [il, s'l_lift, hp, tangent_bundle.proj] with mfld_simps }, + simp only [s'_lift] with mfld_simps at hp, + simp only [il, s'l_lift, hp, total_space.proj] with mfld_simps }, /- Third step: check that the composition of all the maps indeed coincides with the derivative we are looking for -/ have eq_comp : ∀q ∈ s'_lift, tangent_map_within I I' f s q = (Dr.symm ∘ ir ∘ (tangent_map_within I I' (r ∘ f ∘ l.symm) s'l) ∘ (il.symm ∘ Dl)) q, { assume q hq, - simp only [s'_lift, tangent_bundle.proj] with mfld_simps at hq, + simp only [s'_lift] with mfld_simps at hq, have U'q : unique_mdiff_within_at I s' q.1, by { apply U', simp only [hq, s'] with mfld_simps }, have U'lq : unique_mdiff_within_at I s'l (Dl q).1, @@ -362,7 +363,7 @@ begin simp only [hq] with mfld_simps }, rw [this, tangent_map_chart_symm, hDl], { simp only [hq] with mfld_simps, - have : q ∈ (chart_at (model_prod H E) p).source, by simp only [hq] with mfld_simps, + have : q ∈ (chart_at (model_prod H E) p).source, { simp only [hq] with mfld_simps }, exact (chart_at (model_prod H E) p).left_inv this }, { simp only [hq] with mfld_simps } }, have C : tangent_map_within I I' (r ∘ f) s' q @@ -382,14 +383,14 @@ begin { apply is_open.unique_mdiff_within_at _ r.open_source, simp [hq] }, { refine mdifferentiable_at_atlas _ (chart_mem_atlas _ _) _, simp only [hq] with mfld_simps } }, - have : f p.1 = (tangent_map_within I I' f s p).1 := rfl, + have : f p.proj = (tangent_map_within I I' f s p).1 := rfl, rw [A], dsimp [r, Dr], rw [this, tangent_map_chart], { simp only [hq] with mfld_simps, have : tangent_map_within I I' f s' q ∈ (chart_at (model_prod H' E') (tangent_map_within I I' f s p)).source, - by simp only [hq] with mfld_simps, + by { simp only [hq] with mfld_simps }, exact (chart_at (model_prod H' E') (tangent_map_within I I' f s p)).left_inv this }, { simp only [hq] with mfld_simps } }, have E : tangent_map_within I I' f s' q = tangent_map_within I I' f s q, @@ -404,10 +405,10 @@ end derivative is continuous there. -/ theorem cont_mdiff_on.continuous_on_tangent_map_within (hf : cont_mdiff_on I I' n f s) (hmn : 1 ≤ n) (hs : unique_mdiff_on I s) : - continuous_on (tangent_map_within I I' f s) ((tangent_bundle.proj I M) ⁻¹' s) := + continuous_on (tangent_map_within I I' f s) (π (tangent_space I) ⁻¹' s) := begin have : cont_mdiff_on I.tangent I'.tangent 0 (tangent_map_within I I' f s) - ((tangent_bundle.proj I M) ⁻¹' s) := + (π (tangent_space I) ⁻¹' s) := hf.cont_mdiff_on_tangent_map_within hmn hs, exact this.continuous_on end @@ -435,169 +436,10 @@ end end tangent_map -/-! ### Smoothness of the projection in a basic smooth bundle -/ - -namespace basic_smooth_vector_bundle_core - -variables (Z : basic_smooth_vector_bundle_core I M E') - -/-- A version of `cont_mdiff_at_iff_target` when the codomain is the total space of - a `basic_smooth_vector_bundle_core`. The continuity condition in the RHS is weaker. -/ -lemma cont_mdiff_at_iff_target {f : N → Z.to_vector_bundle_core.total_space} - {x : N} {n : ℕ∞} : - cont_mdiff_at J (I.prod 𝓘(𝕜, E')) n f x ↔ continuous_at (bundle.total_space.proj ∘ f) x ∧ - cont_mdiff_at J 𝓘(𝕜, E × E') n (ext_chart_at (I.prod 𝓘(𝕜, E')) (f x) ∘ f) x := -begin - let Z' := Z.to_vector_bundle_core, - rw [cont_mdiff_at_iff_target, and.congr_left_iff], - refine λ hf, ⟨λ h, Z'.continuous_proj.continuous_at.comp h, λ h, _⟩, - exact (Z'.local_triv ⟨chart_at _ (f x).1, chart_mem_atlas _ _⟩) - .continuous_at_of_comp_left h (mem_chart_source _ _) (h.prod hf.continuous_at.snd) -end - -lemma smooth_iff_target {f : N → Z.to_vector_bundle_core.total_space} : - smooth J (I.prod 𝓘(𝕜, E')) f ↔ continuous (bundle.total_space.proj ∘ f) ∧ - ∀ x, smooth_at J 𝓘(𝕜, E × E') (ext_chart_at (I.prod 𝓘(𝕜, E')) (f x) ∘ f) x := -by simp_rw [smooth, smooth_at, cont_mdiff, Z.cont_mdiff_at_iff_target, forall_and_distrib, - continuous_iff_continuous_at] - -lemma cont_mdiff_proj : - cont_mdiff (I.prod 𝓘(𝕜, E')) I n Z.to_vector_bundle_core.proj := -begin - assume x, - rw [cont_mdiff_at, cont_mdiff_within_at_iff'], - refine ⟨Z.to_vector_bundle_core.continuous_proj.continuous_within_at, _⟩, - simp only [(∘), chart_at, chart] with mfld_simps, - apply cont_diff_within_at_fst.congr, - { rintros ⟨a, b⟩ hab, - simp only with mfld_simps at hab, - simp only [hab] with mfld_simps }, - { simp only with mfld_simps } -end - -lemma smooth_proj : - smooth (I.prod 𝓘(𝕜, E')) I Z.to_vector_bundle_core.proj := -cont_mdiff_proj Z - -lemma cont_mdiff_on_proj {s : set (Z.to_vector_bundle_core.total_space)} : - cont_mdiff_on (I.prod 𝓘(𝕜, E')) I n - Z.to_vector_bundle_core.proj s := -Z.cont_mdiff_proj.cont_mdiff_on - -lemma smooth_on_proj {s : set (Z.to_vector_bundle_core.total_space)} : - smooth_on (I.prod 𝓘(𝕜, E')) I Z.to_vector_bundle_core.proj s := -cont_mdiff_on_proj Z - -lemma cont_mdiff_at_proj {p : Z.to_vector_bundle_core.total_space} : - cont_mdiff_at (I.prod 𝓘(𝕜, E')) I n - Z.to_vector_bundle_core.proj p := -Z.cont_mdiff_proj.cont_mdiff_at - -lemma smooth_at_proj {p : Z.to_vector_bundle_core.total_space} : - smooth_at (I.prod 𝓘(𝕜, E')) I Z.to_vector_bundle_core.proj p := -Z.cont_mdiff_at_proj - -lemma cont_mdiff_within_at_proj - {s : set (Z.to_vector_bundle_core.total_space)} - {p : Z.to_vector_bundle_core.total_space} : - cont_mdiff_within_at (I.prod 𝓘(𝕜, E')) I n - Z.to_vector_bundle_core.proj s p := -Z.cont_mdiff_at_proj.cont_mdiff_within_at - -lemma smooth_within_at_proj - {s : set (Z.to_vector_bundle_core.total_space)} - {p : Z.to_vector_bundle_core.total_space} : - smooth_within_at (I.prod 𝓘(𝕜, E')) I - Z.to_vector_bundle_core.proj s p := -Z.cont_mdiff_within_at_proj - -/-- If an element of `E'` is invariant under all coordinate changes, then one can define a -corresponding section of the fiber bundle, which is smooth. This applies in particular to the -zero section of a vector bundle. Another example (not yet defined) would be the identity -section of the endomorphism bundle of a vector bundle. -/ -lemma smooth_const_section (v : E') - (h : ∀ (i j : atlas H M), ∀ x ∈ i.1.source ∩ j.1.source, Z.coord_change i j (i.1 x) v = v) : - smooth I (I.prod 𝓘(𝕜, E')) - (show M → Z.to_vector_bundle_core.total_space, from λ x, ⟨x, v⟩) := -begin - assume x, - rw [cont_mdiff_at, cont_mdiff_within_at_iff'], - split, - { apply continuous.continuous_within_at, - apply fiber_bundle_core.continuous_const_section, - assume i j y hy, - exact h _ _ _ hy }, - { have : cont_diff 𝕜 ⊤ (λ (y : E), (y, v)) := cont_diff_id.prod cont_diff_const, - apply this.cont_diff_within_at.congr, - { assume y hy, - simp only with mfld_simps at hy, - simp only [chart, hy, chart_at, prod.mk.inj_iff, to_vector_bundle_core] - with mfld_simps, - apply h, - simp only [hy, subtype.val_eq_coe] with mfld_simps }, - { simp only [chart, chart_at, prod.mk.inj_iff, to_vector_bundle_core] - with mfld_simps, - apply h, - simp only [subtype.val_eq_coe] with mfld_simps } } -end - -end basic_smooth_vector_bundle_core - -/-! ### Smoothness of the tangent bundle projection -/ - namespace tangent_bundle include Is - -lemma cont_mdiff_proj : - cont_mdiff I.tangent I n (proj I M) := -basic_smooth_vector_bundle_core.cont_mdiff_proj _ - -lemma smooth_proj : smooth I.tangent I (proj I M) := -basic_smooth_vector_bundle_core.smooth_proj _ - -lemma cont_mdiff_on_proj {s : set (tangent_bundle I M)} : - cont_mdiff_on I.tangent I n (proj I M) s := -basic_smooth_vector_bundle_core.cont_mdiff_on_proj _ - -lemma smooth_on_proj {s : set (tangent_bundle I M)} : - smooth_on I.tangent I (proj I M) s := -basic_smooth_vector_bundle_core.smooth_on_proj _ - -lemma cont_mdiff_at_proj {p : tangent_bundle I M} : - cont_mdiff_at I.tangent I n - (proj I M) p := -basic_smooth_vector_bundle_core.cont_mdiff_at_proj _ - -lemma smooth_at_proj {p : tangent_bundle I M} : - smooth_at I.tangent I (proj I M) p := -basic_smooth_vector_bundle_core.smooth_at_proj _ - -lemma cont_mdiff_within_at_proj - {s : set (tangent_bundle I M)} {p : tangent_bundle I M} : - cont_mdiff_within_at I.tangent I n - (proj I M) s p := -basic_smooth_vector_bundle_core.cont_mdiff_within_at_proj _ - -lemma smooth_within_at_proj - {s : set (tangent_bundle I M)} {p : tangent_bundle I M} : - smooth_within_at I.tangent I - (proj I M) s p := -basic_smooth_vector_bundle_core.smooth_within_at_proj _ - variables (I M) -/-- The zero section of the tangent bundle -/ -def zero_section : M → tangent_bundle I M := λ x, ⟨x, 0⟩ -variables {I M} - -lemma smooth_zero_section : smooth I I.tangent (zero_section I M) := -begin - apply basic_smooth_vector_bundle_core.smooth_const_section (tangent_bundle_core I M) 0, - assume i j x hx, - simp only [tangent_bundle_core, continuous_linear_map.map_zero, continuous_linear_map.coe_coe] - with mfld_simps, -end - open bundle /-- The derivative of the zero section of the tangent bundle maps `⟨x, v⟩` to `⟨⟨x, 0⟩, ⟨v, 0⟩⟩`. @@ -615,16 +457,18 @@ may seem. TODO define splittings of vector bundles; state this result invariantly. -/ lemma tangent_map_tangent_bundle_pure (p : tangent_bundle I M) : - tangent_map I I.tangent (tangent_bundle.zero_section I M) p = ⟨⟨p.1, 0⟩, ⟨p.2, 0⟩⟩ := + tangent_map I I.tangent (zero_section (tangent_space I)) p = ⟨⟨p.proj, 0⟩, ⟨p.2, 0⟩⟩ := begin rcases p with ⟨x, v⟩, have N : I.symm ⁻¹' (chart_at H x).target ∈ 𝓝 (I ((chart_at H x) x)), { apply is_open.mem_nhds, apply (local_homeomorph.open_target _).preimage I.continuous_inv_fun, simp only with mfld_simps }, - have A : mdifferentiable_at I I.tangent (λ x, @total_space_mk M (tangent_space I) x 0) x := - tangent_bundle.smooth_zero_section.mdifferentiable_at, - have B : fderiv_within 𝕜 (λ (x_1 : E), (x_1, (0 : E))) (set.range ⇑I) (I ((chart_at H x) x)) v + have A : mdifferentiable_at I I.tangent (λ x, @total_space_mk M (tangent_space I) x 0) x, + { have : smooth I (I.prod 𝓘(𝕜, E)) (zero_section (tangent_space I : M → Type*)) := + bundle.smooth_zero_section 𝕜 (tangent_space I : M → Type*), + exact this.mdifferentiable_at }, + have B : fderiv_within 𝕜 (λ (x' : E), (x', (0 : E))) (set.range ⇑I) (I ((chart_at H x) x)) v = (v, 0), { rw [fderiv_within_eq_fderiv, differentiable_at.fderiv_prod], { simp }, @@ -632,10 +476,9 @@ begin { exact differentiable_at_const _ }, { exact model_with_corners.unique_diff_at_image I }, { exact differentiable_at_id'.prod (differentiable_at_const _) } }, - simp only [tangent_bundle.zero_section, tangent_map, mfderiv, - A, if_pos, chart_at, basic_smooth_vector_bundle_core.chart, - basic_smooth_vector_bundle_core.to_vector_bundle_core, tangent_bundle_core, - function.comp, continuous_linear_map.map_zero] with mfld_simps, + simp only [bundle.zero_section, tangent_map, mfderiv, total_space.proj_mk, A, + if_pos, chart_at, fiber_bundle.charted_space_chart_at, tangent_bundle.trivialization_at_apply, + tangent_bundle_core, function.comp, continuous_linear_map.map_zero] with mfld_simps, rw ← fderiv_within_inter N (I.unique_diff (I ((chart_at H x) x)) (set.mem_range_self _)) at B, rw [← fderiv_within_inter N (I.unique_diff (I ((chart_at H x) x)) (set.mem_range_self _)), ← B], congr' 2, diff --git a/src/geometry/manifold/mfderiv.lean b/src/geometry/manifold/mfderiv.lean index 557a2322c06b3..e80c4ec556fb4 100644 --- a/src/geometry/manifold/mfderiv.lean +++ b/src/geometry/manifold/mfderiv.lean @@ -1,10 +1,9 @@ /- Copyright (c) 2020 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Sébastien Gouëzel +Authors: Sébastien Gouëzel, Floris van Doorn -/ -import geometry.manifold.local_invariant_properties -import geometry.manifold.tangent_bundle +import geometry.manifold.vector_bundle.tangent /-! # The derivative of functions between smooth manifolds @@ -61,9 +60,9 @@ of `f` in these charts. Due to the fact that we are working in a model with corners, with an additional embedding `I` of the model space `H` in the model vector space `E`, the charts taking values in `E` are not the original charts of the manifold, but those ones composed with `I`, called extended charts. We define -`written_in_ext_chart I I' x f` for the function `f` written in the preferred extended charts. Then -the manifold derivative of `f`, at `x`, is just the usual derivative of `written_in_ext_chart I I' x -f`, at the point `(ext_chart_at I x) x`. +`written_in_ext_chart I I' x f` for the function `f` written in the preferred extended charts. Then +the manifold derivative of `f`, at `x`, is just the usual derivative of +`written_in_ext_chart I I' x f`, at the point `(ext_chart_at I x) x`. There is a subtelty with respect to continuity: if the function is not continuous, then the image of a small open set around `x` will not be contained in the source of the preferred chart around @@ -93,7 +92,7 @@ Derivative, manifold -/ noncomputable theory -open_locale classical topology manifold +open_locale classical topology manifold bundle open set @@ -683,16 +682,16 @@ begin exact tangent_map_within_subset (subset_univ _) hs h, end -@[simp, mfld_simps] lemma tangent_map_within_tangent_bundle_proj {p : tangent_bundle I M} : - tangent_bundle.proj I' M' (tangent_map_within I I' f s p) = f (tangent_bundle.proj I M p) := rfl - @[simp, mfld_simps] lemma tangent_map_within_proj {p : tangent_bundle I M} : - (tangent_map_within I I' f s p).1 = f p.1 := rfl + (tangent_map_within I I' f s p).proj = f p.proj := rfl -@[simp, mfld_simps] lemma tangent_map_tangent_bundle_proj {p : tangent_bundle I M} : - tangent_bundle.proj I' M' (tangent_map I I' f p) = f (tangent_bundle.proj I M p) := rfl +@[simp, mfld_simps] lemma tangent_map_within_fst {p : tangent_bundle I M} : + (tangent_map_within I I' f s p).1 = f p.1 := rfl @[simp, mfld_simps] lemma tangent_map_proj {p : tangent_bundle I M} : + (tangent_map I I' f p).proj = f p.proj := rfl + +@[simp, mfld_simps] lemma tangent_map_fst {p : tangent_bundle I M} : (tangent_map I I' f p).1 = f p.1 := rfl omit Is I's @@ -1176,7 +1175,7 @@ end by { ext1 ⟨x, v⟩, simp [tangent_map] } lemma tangent_map_within_id {p : tangent_bundle I M} - (hs : unique_mdiff_within_at I s (tangent_bundle.proj I M p)) : + (hs : unique_mdiff_within_at I s p.proj) : tangent_map_within I I (id : M → M) s p = p := begin simp only [tangent_map_within, id.def], @@ -1496,9 +1495,8 @@ begin rw mdifferentiable_at.mfderiv (mdifferentiable_at_atlas_symm _ (chart_mem_atlas _ _) h), -- a trivial instance is needed after the rewrite, handle it right now. rotate, { apply_instance }, - simp only [continuous_linear_map.coe_coe, basic_smooth_vector_bundle_core.chart, h, - tangent_bundle_core, basic_smooth_vector_bundle_core.to_vector_bundle_core, - chart_at, sigma.mk.inj_iff] with mfld_simps, + simp only [continuous_linear_map.coe_coe, tangent_bundle.chart_at, h, + tangent_bundle_core, chart_at, sigma.mk.inj_iff] with mfld_simps, end end charts @@ -1761,13 +1759,16 @@ begin exact this.unique_diff_on_target_inter _ end +open bundle variables {F : Type*} [normed_add_comm_group F] [normed_space 𝕜 F] -(Z : basic_smooth_vector_bundle_core I M F) + (Z : M → Type*) [topological_space (total_space Z)] [∀ b, topological_space (Z b)] + [∀ b, add_comm_monoid (Z b)] [∀ b, module 𝕜 (Z b)] + [fiber_bundle F Z] [vector_bundle 𝕜 F Z] [smooth_vector_bundle F Z I] -/-- In a smooth fiber bundle constructed from core, the preimage under the projection of a set with +/-- In a smooth fiber bundle, the preimage under the projection of a set with unique differential in the basis also has unique differential. -/ lemma unique_mdiff_on.smooth_bundle_preimage (hs : unique_mdiff_on I s) : - unique_mdiff_on (I.prod (𝓘(𝕜, F))) (Z.to_vector_bundle_core.proj ⁻¹' s) := + unique_mdiff_on (I.prod (𝓘(𝕜, F))) (π Z ⁻¹' s) := begin /- Using a chart (and the fact that unique differentiability is invariant under charts), we reduce the situation to the model space, where we can use the fact that products respect @@ -1776,31 +1777,48 @@ begin replace hp : p.fst ∈ s, by simpa only with mfld_simps using hp, let e₀ := chart_at H p.1, let e := chart_at (model_prod H F) p, + have h2s : ∀ x, x ∈ e.target ∩ e.symm ⁻¹' (π Z ⁻¹' s) ↔ + (x.1 ∈ e₀.target ∧ (e₀.symm) x.1 ∈ (trivialization_at F Z p.1).base_set) ∧ (e₀.symm) x.1 ∈ s, + { intro x, + have A : x ∈ e.target ↔ x.1 ∈ e₀.target ∧ + (e₀.symm) x.1 ∈ (trivialization_at F Z p.1).base_set, + { simp only [e, fiber_bundle.charted_space_chart_at, trivialization.mem_target, + bundle.total_space.proj] with mfld_simps }, + rw [← A, mem_inter_iff, and.congr_right_iff], + intro hx, + simp only [fiber_bundle.charted_space_chart_at_symm_fst p x hx] with mfld_simps }, -- It suffices to prove unique differentiability in a chart suffices h : unique_mdiff_on (I.prod (𝓘(𝕜, F))) - (e.target ∩ e.symm⁻¹' (Z.to_vector_bundle_core.proj ⁻¹' s)), + (e.target ∩ e.symm ⁻¹' (π Z ⁻¹' s)), { have A : unique_mdiff_on (I.prod (𝓘(𝕜, F))) (e.symm.target ∩ - e.symm.symm ⁻¹' (e.target ∩ e.symm⁻¹' (Z.to_vector_bundle_core.proj ⁻¹' s))), + e.symm.symm ⁻¹' (e.target ∩ e.symm⁻¹' (π Z ⁻¹' s))), { apply h.unique_mdiff_on_preimage, exact (mdifferentiable_of_mem_atlas _ (chart_mem_atlas _ _)).symm, apply_instance }, - have : p ∈ e.symm.target ∩ - e.symm.symm ⁻¹' (e.target ∩ e.symm⁻¹' (Z.to_vector_bundle_core.proj ⁻¹' s)), - by simp only [e, hp] with mfld_simps, + have : p ∈ e.symm.target ∩ e.symm.symm ⁻¹' (e.target ∩ e.symm⁻¹' (π Z ⁻¹' s)), + { simp only [e, hp] with mfld_simps }, apply (A _ this).mono, assume q hq, simp only [e, local_homeomorph.left_inv _ hq.1] with mfld_simps at hq, simp only [hq] with mfld_simps }, - -- rewrite the relevant set in the chart as a direct product - have : (λ (p : E × F), (I.symm p.1, p.snd)) ⁻¹' e.target ∩ - (λ (p : E × F), (I.symm p.1, p.snd)) ⁻¹' (e.symm ⁻¹' (sigma.fst ⁻¹' s)) ∩ - (range I ×ˢ univ) - = (I.symm ⁻¹' (e₀.target ∩ e₀.symm⁻¹' s) ∩ range I) ×ˢ univ, - by mfld_set_tac, assume q hq, - replace hq : q.1 ∈ (chart_at H p.1).target ∧ ((chart_at H p.1).symm : H → M) q.1 ∈ s, - by simpa only with mfld_simps using hq, - simp only [unique_mdiff_within_at, model_with_corners.prod, preimage_inter, this] with mfld_simps, + simp only [unique_mdiff_within_at, model_with_corners.prod, -preimage_inter] with mfld_simps, + have : 𝓝[(I.symm ⁻¹' (e₀.target ∩ e₀.symm⁻¹' s) ∩ range I) ×ˢ univ] (I q.1, q.2) ≤ + 𝓝[(λ (p : E × F), (I.symm p.1, p.snd)) ⁻¹' (e.target ∩ e.symm ⁻¹' (π Z ⁻¹' s)) ∩ + (range I ×ˢ univ)] (I q.1, q.2), + { rw [nhds_within_le_iff, mem_nhds_within], + refine ⟨(λ (p : E × F), (I.symm p.1, p.snd)) ⁻¹' e.target, _, _, _⟩, + { exact e.open_target.preimage (I.continuous_symm.prod_map continuous_id) }, + { simp only [prod.mk.eta] with mfld_simps at hq, + simp only [prod.mk.eta, hq] with mfld_simps }, + rintro x hx, + simp only with mfld_simps at hx, + have h2x := hx, + simp only [e, fiber_bundle.charted_space_chart_at, trivialization.mem_target] + with mfld_simps at h2x, + simp only [h2s, hx, h2x, -preimage_inter] with mfld_simps }, + refine unique_diff_within_at.mono_nhds _ this, + rw [h2s] at hq, -- apply unique differentiability of products to conclude apply unique_diff_on.prod _ unique_diff_on_univ, { simp only [hq] with mfld_simps }, @@ -1815,7 +1833,7 @@ begin end lemma unique_mdiff_on.tangent_bundle_proj_preimage (hs : unique_mdiff_on I s): - unique_mdiff_on I.tangent ((tangent_bundle.proj I M) ⁻¹' s) := + unique_mdiff_on I.tangent (π (tangent_space I) ⁻¹' s) := hs.smooth_bundle_preimage _ end unique_mdiff diff --git a/src/geometry/manifold/tangent_bundle.lean b/src/geometry/manifold/tangent_bundle.lean deleted file mode 100644 index a761b141522da..0000000000000 --- a/src/geometry/manifold/tangent_bundle.lean +++ /dev/null @@ -1,651 +0,0 @@ -/- -Copyright (c) 2019 Sébastien Gouëzel. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Sébastien Gouëzel --/ -import topology.vector_bundle.basic -import geometry.manifold.smooth_manifold_with_corners -import data.set.prod - -/-! -# Basic smooth bundles - -In general, a smooth bundle is a bundle over a smooth manifold, whose fiber is a manifold, and -for which the coordinate changes are smooth. In this definition, there are charts involved at -several places: in the manifold structure of the base, in the manifold structure of the fibers, and -in the local trivializations. This makes it a complicated object in general. There is however a -specific situation where things are much simpler: when the fiber is a vector space (no need for -charts for the fibers), and when the local trivializations of the bundle and the charts of the base -coincide. Then everything is expressed in terms of the charts of the base, making for a much -simpler overall structure, which is easier to manipulate formally. - -Most vector bundles that naturally occur in differential geometry are of this form: -the tangent bundle, the cotangent bundle, differential forms (used to define de Rham cohomology) -and the bundle of Riemannian metrics. Therefore, it is worth defining a specific constructor for -this kind of bundle, that we call basic smooth bundles. - -A basic smooth bundle is thus a smooth bundle over a smooth manifold whose fiber is a vector space, -and which is trivial in the coordinate charts of the base. (We recall that in our notion of manifold -there is a distinguished atlas, which does not need to be maximal: we require the triviality above -this specific atlas). It can be constructed from a basic smooth bundled core, defined below, -specifying the changes in the fiber when one goes from one coordinate chart to another one. - -## Main definitions - -* `basic_smooth_vector_bundle_core I M F`: assuming that `M` is a smooth manifold over the model - with corners `I` on `(𝕜, E, H)`, and `F` is a normed vector space over `𝕜`, this structure - registers, for each pair of charts of `M`, a linear change of coordinates on `F` depending - smoothly on the base point. This is the core structure from which one will build a smooth vector - bundle with fiber `F` over `M`. - -Let `Z` be a basic smooth bundle core over `M` with fiber `F`. We define -`Z.to_vector_bundle_core`, the (topological) vector bundle core associated to `Z`. From -it, we get a space `Z.to_vector_bundle_core.total_space` (which as a Type is just -`Σ (x : M), F`), with the fiber bundle topology. It inherits a manifold structure (where the -charts are in bijection with the charts of the basis). We show that this manifold is smooth. - -Then we use this machinery to construct the tangent bundle of a smooth manifold. - -* `tangent_bundle_core I M`: the basic smooth bundle core associated to a smooth manifold `M` over - a model with corners `I`. -* `tangent_bundle I M` : the total space of `tangent_bundle_core I M`. It is itself a - smooth manifold over the model with corners `I.tangent`, the product of `I` and the trivial model - with corners on `E`. -* `tangent_space I x` : the tangent space to `M` at `x` -* `tangent_bundle.proj I M`: the projection from the tangent bundle to the base manifold - -## Implementation notes - -We register the vector space structure on the fibers of the tangent bundle, but we do not register -the normed space structure coming from that of `F` (as it is not canonical, and we also want to -keep the possibility to add a Riemannian structure on the manifold later on without having two -competing normed space instances on the tangent spaces). - -We require `F` to be a normed space, and not just a topological vector space, as we want to talk -about smooth functions on `F`. The notion of derivative requires a norm to be defined. - -## TODO -construct the cotangent bundle, and the bundles of differential forms. They should follow -functorially from the description of the tangent bundle as a basic smooth bundle. - -## Tags -Smooth fiber bundle, vector bundle, tangent space, tangent bundle --/ -noncomputable theory - -universe u - -open topological_space set -open_locale manifold topology - -/-- Core structure used to create a smooth bundle above `M` (a manifold over the model with -corner `I`) with fiber the normed vector space `F` over `𝕜`, which is trivial in the chart domains -of `M`. This structure registers the changes in the fibers when one changes coordinate charts in the -base. We require the change of coordinates of the fibers to be linear, so that the resulting bundle -is a vector bundle. -/ -structure basic_smooth_vector_bundle_core {𝕜 : Type*} [nontrivially_normed_field 𝕜] -{E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] -{H : Type*} [topological_space H] (I : model_with_corners 𝕜 E H) -(M : Type*) [topological_space M] [charted_space H M] [smooth_manifold_with_corners I M] -(F : Type*) [normed_add_comm_group F] [normed_space 𝕜 F] := -(coord_change : atlas H M → atlas H M → H → (F →L[𝕜] F)) -(coord_change_self : ∀ i : atlas H M, ∀ x ∈ i.1.target, ∀ v, coord_change i i x v = v) -(coord_change_comp : ∀ i j k : atlas H M, - ∀ x ∈ ((i.1.symm.trans j.1).trans (j.1.symm.trans k.1)).source, ∀ v, - (coord_change j k ((i.1.symm.trans j.1) x)) (coord_change i j x v) = coord_change i k x v) -(coord_change_smooth_clm : ∀ i j : atlas H M, - cont_diff_on 𝕜 ∞ ((coord_change i j) ∘ I.symm) (I '' (i.1.symm.trans j.1).source)) - -/-- The trivial basic smooth bundle core, in which all the changes of coordinates are the -identity. -/ -def trivial_basic_smooth_vector_bundle_core {𝕜 : Type*} [nontrivially_normed_field 𝕜] -{E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] -{H : Type*} [topological_space H] (I : model_with_corners 𝕜 E H) -(M : Type*) [topological_space M] [charted_space H M] [smooth_manifold_with_corners I M] -(F : Type*) [normed_add_comm_group F] [normed_space 𝕜 F] : basic_smooth_vector_bundle_core I M F := -{ coord_change := λ i j x, continuous_linear_map.id 𝕜 F, - coord_change_self := λ i x hx v, rfl, - coord_change_comp := λ i j k x hx v, rfl, - coord_change_smooth_clm := λ i j, by { dsimp, exact cont_diff_on_const } } - -namespace basic_smooth_vector_bundle_core - -variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] -{E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] -{H : Type*} [topological_space H] {I : model_with_corners 𝕜 E H} -{M : Type*} [topological_space M] [charted_space H M] [smooth_manifold_with_corners I M] -{F : Type*} [normed_add_comm_group F] [normed_space 𝕜 F] -(Z : basic_smooth_vector_bundle_core I M F) - -instance : inhabited (basic_smooth_vector_bundle_core I M F) := -⟨trivial_basic_smooth_vector_bundle_core I M F⟩ - -/-- A reformulation of `coord_change_comp`, formulated in terms of a point in `M`. -The conditions on this point a significantly simpler than in `coord_change_comp`. -/ -lemma coord_change_comp' {i j k : atlas H M} {x : M} - (hi : x ∈ i.1.source) (hj : x ∈ j.1.source) (hk : x ∈ k.1.source) (v : F) : - Z.coord_change j k (j x) (Z.coord_change i j (i x) v) = Z.coord_change i k (i x) v := -begin - rw [show j x = _, by rw [← i.1.left_inv hi]], - apply Z.coord_change_comp, - simp only [hi, hj, hk] with mfld_simps -end - -/-- A reformulation of `coord_change_self`, formulated in terms of a point in `M`. -/ -lemma coord_change_self' {i : atlas H M} {x : M} (hi : x ∈ i.1.source) (v : F) : - Z.coord_change i i (i x) v = v := -Z.coord_change_self i (i x) (i.1.maps_to hi) v - -/-- `Z.coord_change j i` is a partial inverse of `Z.coord_change i j`. -/ -lemma coord_change_comp_eq_self (i j : atlas H M) {x : H} - (hx : x ∈ (i.1.symm.trans j.1).source) (v : F) : - Z.coord_change j i (i.1.symm.trans j.1 x) (Z.coord_change i j x v) = v := -begin - rw [Z.coord_change_comp i j i x _ v, Z.coord_change_self _ _ hx.1], - simp only with mfld_simps at hx, - simp only [hx.1, hx.2] with mfld_simps -end - -/-- `Z.coord_change j i` is a partial inverse of `Z.coord_change i j`, -formulated in terms of a point in `M`. -/ -lemma coord_change_comp_eq_self' {i j : atlas H M} {x : M} - (hi : x ∈ i.1.source) (hj : x ∈ j.1.source) (v : F) : - Z.coord_change j i (j x) (Z.coord_change i j (i x) v) = v := -by rw [Z.coord_change_comp' hi hj hi v, Z.coord_change_self' hi] - -lemma coord_change_continuous (i j : atlas H M) : - continuous_on (Z.coord_change i j) (i.1.symm.trans j.1).source := -begin - assume x hx, - apply (((Z.coord_change_smooth_clm i j).continuous_on.continuous_within_at - (mem_image_of_mem I hx)).comp I.continuous_within_at _).congr, - { assume y hy, - simp only with mfld_simps }, - { simp only with mfld_simps }, - { exact maps_to_image I _ }, -end - -lemma coord_change_smooth (i j : atlas H M) : - cont_diff_on 𝕜 ∞ (λ p : E × F, Z.coord_change i j (I.symm p.1) p.2) - ((I '' (i.1.symm.trans j.1).source) ×ˢ univ) := -begin - have A : cont_diff 𝕜 ∞ (λ p : (F →L[𝕜] F) × F, p.1 p.2), - { apply is_bounded_bilinear_map.cont_diff, - exact is_bounded_bilinear_map_apply }, - have B : cont_diff_on 𝕜 ∞ (λ (p : E × F), (Z.coord_change i j (I.symm p.1), p.snd)) - ((I '' (i.1.symm.trans j.1).source) ×ˢ univ), - { apply cont_diff_on.prod _ _, - { exact (Z.coord_change_smooth_clm i j).comp cont_diff_fst.cont_diff_on - (prod_subset_preimage_fst _ _) }, - { exact is_bounded_linear_map.snd.cont_diff.cont_diff_on } }, - exact A.comp_cont_diff_on B, -end - -/-- Vector bundle core associated to a basic smooth bundle core -/ -@[simps coord_change index_at] -def to_vector_bundle_core : vector_bundle_core 𝕜 M F (atlas H M) := -{ base_set := λ i, i.1.source, - is_open_base_set := λ i, i.1.open_source, - index_at := achart H, - mem_base_set_at := λ x, mem_chart_source H x, - coord_change := λ i j x, Z.coord_change i j (i.1 x), - coord_change_self := λ i x hx v, Z.coord_change_self i (i.1 x) (i.1.map_source hx) v, - coord_change_comp := λ i j k x ⟨⟨hx1, hx2⟩, hx3⟩ v, begin - have := Z.coord_change_comp i j k (i.1 x) _ v, - convert this using 2, - { simp only [hx1] with mfld_simps }, - { simp only [hx1, hx2, hx3] with mfld_simps } - end, - continuous_on_coord_change := λ i j, begin - refine ((Z.coord_change_continuous i j).comp' i.1.continuous_on).mono _, - rintros p ⟨hp₁, hp₂⟩, - refine ⟨hp₁, i.1.maps_to hp₁, _⟩, - simp only [i.1.left_inv hp₁, hp₂] with mfld_simps - end } - -@[simp, mfld_simps] lemma base_set (i : atlas H M) : - (Z.to_vector_bundle_core.local_triv i).base_set = i.1.source := rfl - -@[simp, mfld_simps] lemma target (i : atlas H M) : - (Z.to_vector_bundle_core.local_triv i).target = i.1.source ×ˢ univ := rfl - -/-- Local chart for the total space of a basic smooth bundle -/ -def chart {e : local_homeomorph M H} (he : e ∈ atlas H M) : - local_homeomorph (Z.to_vector_bundle_core.total_space) (model_prod H F) := -(Z.to_vector_bundle_core.local_triv ⟨e, he⟩).to_local_homeomorph.trans - (local_homeomorph.prod e (local_homeomorph.refl F)) - -lemma chart_apply {x : M} (z : Z.to_vector_bundle_core.total_space) : - Z.chart (chart_mem_atlas H x) z = (chart_at H x z.proj, - Z.coord_change (achart H z.proj) (achart H x) (achart H z.proj z.proj) z.2) := -rfl - -@[simp, mfld_simps] lemma chart_source (e : local_homeomorph M H) (he : e ∈ atlas H M) : - (Z.chart he).source = Z.to_vector_bundle_core.proj ⁻¹' e.source := -by { simp only [chart, mem_prod], mfld_set_tac } - -@[simp, mfld_simps] lemma chart_target (e : local_homeomorph M H) (he : e ∈ atlas H M) : - (Z.chart he).target = e.target ×ˢ univ := -by { simp only [chart], mfld_set_tac } - -/-- The total space of a basic smooth bundle is endowed with a charted space structure, where the -charts are in bijection with the charts of the basis. -/ -@[simps chart_at (lemmas_only)] -instance to_charted_space : - charted_space (model_prod H F) Z.to_vector_bundle_core.total_space := -{ atlas := ⋃(e : local_homeomorph M H) (he : e ∈ atlas H M), {Z.chart he}, - chart_at := λ p, Z.chart (chart_mem_atlas H p.1), - mem_chart_source := λ p, by simp [mem_chart_source], - chart_mem_atlas := λ p, begin - simp only [mem_Union, mem_singleton_iff, chart_mem_atlas], - exact ⟨chart_at H p.1, chart_mem_atlas H p.1, rfl⟩ - end } - -lemma mem_atlas_iff - (f : local_homeomorph Z.to_vector_bundle_core.total_space (model_prod H F)) : - f ∈ atlas (model_prod H F) Z.to_vector_bundle_core.total_space ↔ - ∃(e : local_homeomorph M H) (he : e ∈ atlas H M), f = Z.chart he := -by simp only [atlas, mem_Union, mem_singleton_iff] - -@[simp, mfld_simps] lemma mem_chart_source_iff - (p q : Z.to_vector_bundle_core.total_space) : - p ∈ (chart_at (model_prod H F) q).source ↔ p.1 ∈ (chart_at H q.1).source := -by simp only [chart_at] with mfld_simps - -@[simp, mfld_simps] lemma mem_chart_target_iff - (p : H × F) (q : Z.to_vector_bundle_core.total_space) : - p ∈ (chart_at (model_prod H F) q).target ↔ p.1 ∈ (chart_at H q.1).target := -by simp only [chart_at] with mfld_simps - -@[simp, mfld_simps] lemma coe_chart_at_fst (p q : Z.to_vector_bundle_core.total_space) : - ((chart_at (model_prod H F) q) p).1 = chart_at H q.1 p.1 := rfl - -@[simp, mfld_simps] lemma coe_chart_at_symm_fst - (p : H × F) (q : Z.to_vector_bundle_core.total_space) : - ((chart_at (model_prod H F) q).symm p).1 = ((chart_at H q.1).symm : H → M) p.1 := rfl - -/-- Smooth manifold structure on the total space of a basic smooth bundle -/ -instance to_smooth_manifold : - smooth_manifold_with_corners (I.prod (𝓘(𝕜, F))) Z.to_vector_bundle_core.total_space := -begin - /- We have to check that the charts belong to the smooth groupoid, i.e., they are smooth on their - source, and their inverses are smooth on the target. Since both objects are of the same kind, it - suffices to prove the first statement in A below, and then glue back the pieces at the end. -/ - let J := model_with_corners.to_local_equiv (I.prod (𝓘(𝕜, F))), - have A : ∀ (e e' : local_homeomorph M H) (he : e ∈ atlas H M) (he' : e' ∈ atlas H M), - cont_diff_on 𝕜 ∞ - (J ∘ ((Z.chart he).symm.trans (Z.chart he')) ∘ J.symm) - (J.symm ⁻¹' ((Z.chart he).symm.trans (Z.chart he')).source ∩ range J), - { assume e e' he he', - have : J.symm ⁻¹' ((chart Z he).symm.trans (chart Z he')).source ∩ range J = - (I.symm ⁻¹' (e.symm.trans e').source ∩ range I) ×ˢ univ, - by { simp only [J, chart, model_with_corners.prod], mfld_set_tac }, - rw this, - -- check separately that the two components of the coordinate change are smooth - apply cont_diff_on.prod, - show cont_diff_on 𝕜 ∞ (λ (p : E × F), (I ∘ e' ∘ e.symm ∘ I.symm) p.1) - ((I.symm ⁻¹' (e.symm.trans e').source ∩ range I) ×ˢ univ), - { -- the coordinate change on the base is just a coordinate change for `M`, smooth since - -- `M` is smooth - have A : cont_diff_on 𝕜 ∞ (I ∘ (e.symm.trans e') ∘ I.symm) - (I.symm ⁻¹' (e.symm.trans e').source ∩ range I) := - (has_groupoid.compatible (cont_diff_groupoid ∞ I) he he').1, - have B : cont_diff_on 𝕜 ∞ (λ p : E × F, p.1) - ((I.symm ⁻¹' (e.symm.trans e').source ∩ range I) ×ˢ univ) := - cont_diff_fst.cont_diff_on, - exact cont_diff_on.comp A B (prod_subset_preimage_fst _ _) }, - show cont_diff_on 𝕜 ∞ (λ (p : E × F), - Z.coord_change ⟨chart_at H (e.symm (I.symm p.1)), _⟩ ⟨e', he'⟩ - ((chart_at H (e.symm (I.symm p.1)) : M → H) (e.symm (I.symm p.1))) - (Z.coord_change ⟨e, he⟩ ⟨chart_at H (e.symm (I.symm p.1)), _⟩ - (e (e.symm (I.symm p.1))) p.2)) - ((I.symm ⁻¹' (e.symm.trans e').source ∩ range I) ×ˢ univ), - { /- The coordinate change in the fiber is more complicated as its definition involves the - reference chart chosen at each point. However, it appears with its inverse, so using the - cocycle property one can get rid of it, and then conclude using the smoothness of the - cocycle as given in the definition of basic smooth bundles. -/ - have := Z.coord_change_smooth ⟨e, he⟩ ⟨e', he'⟩, - rw I.image_eq at this, - apply cont_diff_on.congr this, - rintros ⟨x, v⟩ hx, - simp only with mfld_simps at hx, - let f := chart_at H (e.symm (I.symm x)), - have A : I.symm x ∈ ((e.symm.trans f).trans (f.symm.trans e')).source, - by simp only [hx.1.1, hx.1.2] with mfld_simps, - rw e.right_inv hx.1.1, - have := Z.coord_change_comp ⟨e, he⟩ ⟨f, chart_mem_atlas _ _⟩ ⟨e', he'⟩ (I.symm x) A v, - simpa only [] using this } }, - refine @smooth_manifold_with_corners.mk _ _ _ _ _ _ _ _ _ _ _ ⟨_⟩, - assume e₀ e₀' he₀ he₀', - rcases (Z.mem_atlas_iff _).1 he₀ with ⟨e, he, rfl⟩, - rcases (Z.mem_atlas_iff _).1 he₀' with ⟨e', he', rfl⟩, - rw [cont_diff_groupoid, mem_groupoid_of_pregroupoid], - exact ⟨A e e' he he', A e' e he' he⟩ -end - -end basic_smooth_vector_bundle_core - -section tangent_bundle - -variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] -{E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] -{H : Type*} [topological_space H] (I : model_with_corners 𝕜 E H) -(M : Type*) [topological_space M] [charted_space H M] [smooth_manifold_with_corners I M] - -/-- Basic smooth bundle core version of the tangent bundle of a smooth manifold `M` modelled over a -model with corners `I` on `(E, H)`. The fibers are equal to `E`, and the coordinate change in the -fiber corresponds to the derivative of the coordinate change in `M`. -/ -@[simps] def tangent_bundle_core : basic_smooth_vector_bundle_core I M E := -{ coord_change := λ i j x, (fderiv_within 𝕜 (I ∘ j.1 ∘ i.1.symm ∘ I.symm) (range I) (I x)), - coord_change_smooth_clm := λ i j, - begin - rw I.image_eq, - have A : cont_diff_on 𝕜 ∞ - (I ∘ (i.1.symm.trans j.1) ∘ I.symm) - (I.symm ⁻¹' (i.1.symm.trans j.1).source ∩ range I) := - (has_groupoid.compatible (cont_diff_groupoid ∞ I) i.2 j.2).1, - have B : unique_diff_on 𝕜 (I.symm ⁻¹' (i.1.symm.trans j.1).source ∩ range I) := - I.unique_diff_preimage_source, - have C : cont_diff_on 𝕜 ∞ - (λ (p : E × E), (fderiv_within 𝕜 (I ∘ j.1 ∘ i.1.symm ∘ I.symm) - (I.symm ⁻¹' (i.1.symm.trans j.1).source ∩ range I) p.1 : E → E) p.2) - ((I.symm ⁻¹' (i.1.symm.trans j.1).source ∩ range I) ×ˢ univ) := - cont_diff_on_fderiv_within_apply A B le_top, - have D : ∀ x ∈ (I.symm ⁻¹' (i.1.symm.trans j.1).source ∩ range I), - fderiv_within 𝕜 (I ∘ j.1 ∘ i.1.symm ∘ I.symm) - (range I) x = - fderiv_within 𝕜 (I ∘ j.1 ∘ i.1.symm ∘ I.symm) - (I.symm ⁻¹' (i.1.symm.trans j.1).source ∩ range I) x, - { assume x hx, - have N : I.symm ⁻¹' (i.1.symm.trans j.1).source ∈ nhds x := - I.continuous_symm.continuous_at.preimage_mem_nhds - (is_open.mem_nhds (local_homeomorph.open_source _) hx.1), - symmetry, - rw inter_comm, - exact fderiv_within_inter N (I.unique_diff _ hx.2) }, - apply (A.fderiv_within B le_top).congr, - assume x hx, - simp only with mfld_simps at hx, - simp only [hx, D] with mfld_simps, - end, - coord_change_self := λ i x hx v, begin - /- Locally, a self-change of coordinate is just the identity, thus its derivative is the - identity. One just needs to write this carefully, paying attention to the sets where the - functions are defined. -/ - have A : I.symm ⁻¹' (i.1.symm.trans i.1).source ∩ range I ∈ 𝓝[range I] (I x), - { rw inter_comm, - apply inter_mem_nhds_within, - apply I.continuous_symm.continuous_at.preimage_mem_nhds - (is_open.mem_nhds (local_homeomorph.open_source _) _), - simp only [hx, i.1.map_target] with mfld_simps }, - have B : ∀ᶠ y in 𝓝[range I] (I x), - (I ∘ i.1 ∘ i.1.symm ∘ I.symm) y = (id : E → E) y, - { filter_upwards [A] with _ hy, - rw ← I.image_eq at hy, - rcases hy with ⟨z, hz⟩, - simp only with mfld_simps at hz, - simp only [hz.2.symm, hz.1] with mfld_simps, }, - have C : fderiv_within 𝕜 (I ∘ i.1 ∘ i.1.symm ∘ I.symm) (range I) (I x) = - fderiv_within 𝕜 (id : E → E) (range I) (I x) := - filter.eventually_eq.fderiv_within_eq I.unique_diff_at_image B - (by simp only [hx] with mfld_simps), - rw fderiv_within_id I.unique_diff_at_image at C, - rw C, - refl - end, - coord_change_comp := λ i j u x hx, begin - /- The cocycle property is just the fact that the derivative of a composition is the product of - the derivatives. One needs however to check that all the functions one considers are smooth, and - to pay attention to the domains where these functions are defined, making this proof a little - bit cumbersome although there is nothing complicated here. -/ - have M : I x ∈ - (I.symm ⁻¹' ((i.1.symm.trans j.1).trans (j.1.symm.trans u.1)).source ∩ range I) := - ⟨by simpa only [mem_preimage, model_with_corners.left_inv] using hx, mem_range_self _⟩, - have U : unique_diff_within_at 𝕜 - (I.symm ⁻¹' ((i.1.symm.trans j.1).trans (j.1.symm.trans u.1)).source ∩ range I) (I x) := - I.unique_diff_preimage_source _ M, - have A : fderiv_within 𝕜 ((I ∘ u.1 ∘ j.1.symm ∘ I.symm) ∘ (I ∘ j.1 ∘ i.1.symm ∘ I.symm)) - (I.symm ⁻¹' ((i.1.symm.trans j.1).trans (j.1.symm.trans u.1)).source ∩ range I) - (I x) - = (fderiv_within 𝕜 (I ∘ u.1 ∘ j.1.symm ∘ I.symm) - (I.symm ⁻¹' (j.1.symm.trans u.1).source ∩ range I) - ((I ∘ j.1 ∘ i.1.symm ∘ I.symm) (I x))).comp - (fderiv_within 𝕜 (I ∘ j.1 ∘ i.1.symm ∘ I.symm) - (I.symm ⁻¹' ((i.1.symm.trans j.1).trans (j.1.symm.trans u.1)).source ∩ range I) - (I x)), - { apply fderiv_within.comp _ _ _ _ U, - show differentiable_within_at 𝕜 (I ∘ j.1 ∘ i.1.symm ∘ I.symm) - (I.symm ⁻¹' ((i.1.symm.trans j.1).trans (j.1.symm.trans u.1)).source ∩ range I) - (I x), - { have A : cont_diff_on 𝕜 ∞ - (I ∘ (i.1.symm.trans j.1) ∘ I.symm) - (I.symm ⁻¹' (i.1.symm.trans j.1).source ∩ range I) := - (has_groupoid.compatible (cont_diff_groupoid ∞ I) i.2 j.2).1, - have B : differentiable_on 𝕜 (I ∘ j.1 ∘ i.1.symm ∘ I.symm) - (I.symm ⁻¹' ((i.1.symm.trans j.1).trans (j.1.symm.trans u.1)).source ∩ range I), - { apply (A.differentiable_on le_top).mono, - have : ((i.1.symm.trans j.1).trans (j.1.symm.trans u.1)).source ⊆ - (i.1.symm.trans j.1).source := inter_subset_left _ _, - exact inter_subset_inter (preimage_mono this) (subset.refl (range I)) }, - apply B, - simpa only [] with mfld_simps using hx }, - show differentiable_within_at 𝕜 (I ∘ u.1 ∘ j.1.symm ∘ I.symm) - (I.symm ⁻¹' (j.1.symm.trans u.1).source ∩ range I) - ((I ∘ j.1 ∘ i.1.symm ∘ I.symm) (I x)), - { have A : cont_diff_on 𝕜 ∞ - (I ∘ (j.1.symm.trans u.1) ∘ I.symm) - (I.symm ⁻¹' (j.1.symm.trans u.1).source ∩ range I) := - (has_groupoid.compatible (cont_diff_groupoid ∞ I) j.2 u.2).1, - apply A.differentiable_on le_top, - rw [local_homeomorph.trans_source] at hx, - simp only with mfld_simps, - exact hx.2 }, - show (I.symm ⁻¹' ((i.1.symm.trans j.1).trans (j.1.symm.trans u.1)).source ∩ range I) - ⊆ (I ∘ j.1 ∘ i.1.symm ∘ I.symm) ⁻¹' (I.symm ⁻¹' (j.1.symm.trans u.1).source ∩ range I), - { assume y hy, - simp only with mfld_simps at hy, - rw [local_homeomorph.left_inv] at hy, - { simp only [hy] with mfld_simps }, - { exact hy.1.1.2 } } }, - have B : fderiv_within 𝕜 ((I ∘ u.1 ∘ j.1.symm ∘ I.symm) - ∘ (I ∘ j.1 ∘ i.1.symm ∘ I.symm)) - (I.symm ⁻¹' ((i.1.symm.trans j.1).trans (j.1.symm.trans u.1)).source ∩ range I) - (I x) - = fderiv_within 𝕜 (I ∘ u.1 ∘ i.1.symm ∘ I.symm) - (I.symm ⁻¹' ((i.1.symm.trans j.1).trans (j.1.symm.trans u.1)).source ∩ range I) - (I x), - { have E : - ∀ y ∈ (I.symm ⁻¹' ((i.1.symm.trans j.1).trans (j.1.symm.trans u.1)).source ∩ range I), - ((I ∘ u.1 ∘ j.1.symm ∘ I.symm) ∘ (I ∘ j.1 ∘ i.1.symm ∘ I.symm)) y = - (I ∘ u.1 ∘ i.1.symm ∘ I.symm) y, - { assume y hy, - simp only [function.comp_app, model_with_corners.left_inv], - rw [j.1.left_inv], - exact hy.1.1.2 }, - exact fderiv_within_congr U E (E _ M) }, - have C : fderiv_within 𝕜 (I ∘ u.1 ∘ i.1.symm ∘ I.symm) - (I.symm ⁻¹' ((i.1.symm.trans j.1).trans (j.1.symm.trans u.1)).source ∩ range I) - (I x) = - fderiv_within 𝕜 (I ∘ u.1 ∘ i.1.symm ∘ I.symm) - (range I) (I x), - { rw inter_comm, - apply fderiv_within_inter _ I.unique_diff_at_image, - apply I.continuous_symm.continuous_at.preimage_mem_nhds - (is_open.mem_nhds (local_homeomorph.open_source _) _), - simpa only [model_with_corners.left_inv] using hx }, - have D : fderiv_within 𝕜 (I ∘ u.1 ∘ j.1.symm ∘ I.symm) - (I.symm ⁻¹' (j.1.symm.trans u.1).source ∩ range I) ((I ∘ j.1 ∘ i.1.symm ∘ I.symm) (I x)) = - fderiv_within 𝕜 (I ∘ u.1 ∘ j.1.symm ∘ I.symm) (range I) ((I ∘ j.1 ∘ i.1.symm ∘ I.symm) (I x)), - { rw inter_comm, - apply fderiv_within_inter _ I.unique_diff_at_image, - apply I.continuous_symm.continuous_at.preimage_mem_nhds - (is_open.mem_nhds (local_homeomorph.open_source _) _), - rw [local_homeomorph.trans_source] at hx, - simp only with mfld_simps, - exact hx.2 }, - have E : fderiv_within 𝕜 (I ∘ j.1 ∘ i.1.symm ∘ I.symm) - (I.symm ⁻¹' ((i.1.symm.trans j.1).trans (j.1.symm.trans u.1)).source ∩ range I) - (I x) = - fderiv_within 𝕜 (I ∘ j.1 ∘ i.1.symm ∘ I.symm) (range I) (I x), - { rw inter_comm, - apply fderiv_within_inter _ I.unique_diff_at_image, - apply I.continuous_symm.continuous_at.preimage_mem_nhds - (is_open.mem_nhds (local_homeomorph.open_source _) _), - simpa only [model_with_corners.left_inv] using hx }, - rw [B, C, D, E] at A, - simp only [A, continuous_linear_map.coe_comp'] with mfld_simps, - end } - -variable {M} -include I - -/-- The tangent space at a point of the manifold `M`. It is just `E`. We could use instead -`(tangent_bundle_core I M).to_vector_bundle_core.fiber x`, but we use `E` to help the -kernel. --/ -@[nolint unused_arguments] -def tangent_space (x : M) : Type* := E - -omit I -variable (M) - -/-- The tangent bundle to a smooth manifold, as a Sigma type. Defined in terms of -`bundle.total_space` to be able to put a suitable topology on it. -/ -@[nolint has_nonempty_instance, reducible] -- is empty if the base manifold is empty -def tangent_bundle := bundle.total_space (tangent_space I : M → Type*) - -local notation `TM` := tangent_bundle I M - -/-- The projection from the tangent bundle of a smooth manifold to the manifold. As the tangent -bundle is represented internally as a sigma type, the notation `p.1` also works for the projection -of the point `p`. -/ -def tangent_bundle.proj : TM → M := -λ p, p.1 - -variable {M} - -@[simp, mfld_simps] lemma tangent_bundle.proj_apply (x : M) (v : tangent_space I x) : - tangent_bundle.proj I M ⟨x, v⟩ = x := -rfl - -section tangent_bundle_instances - -/- In general, the definition of tangent_bundle and tangent_space are not reducible, so that type -class inference does not pick wrong instances. In this section, we record the right instances for -them, noting in particular that the tangent bundle is a smooth manifold. -/ - -section -local attribute [reducible] tangent_space - -variables {M} (x : M) - -instance : topological_space (tangent_space I x) := by apply_instance -instance : add_comm_group (tangent_space I x) := by apply_instance -instance : topological_add_group (tangent_space I x) := by apply_instance -instance : module 𝕜 (tangent_space I x) := by apply_instance -instance : inhabited (tangent_space I x) := ⟨0⟩ - -end - -variable (M) - -instance : topological_space TM := -(tangent_bundle_core I M).to_vector_bundle_core.to_topological_space - -instance : charted_space (model_prod H E) TM := -(tangent_bundle_core I M).to_charted_space - -instance : smooth_manifold_with_corners I.tangent TM := -(tangent_bundle_core I M).to_smooth_manifold - -instance : fiber_bundle E (tangent_space I : M → Type*) := -(tangent_bundle_core I M).to_vector_bundle_core.fiber_bundle - -instance : vector_bundle 𝕜 E (tangent_space I : M → Type*) := -(tangent_bundle_core I M).to_vector_bundle_core.vector_bundle - -end tangent_bundle_instances - -variable (M) - -/-- The tangent bundle projection on the basis is a continuous map. -/ -lemma tangent_bundle_proj_continuous : continuous (tangent_bundle.proj I M) := -((tangent_bundle_core I M).to_vector_bundle_core).continuous_proj - -/-- The tangent bundle projection on the basis is an open map. -/ -lemma tangent_bundle_proj_open : is_open_map (tangent_bundle.proj I M) := -((tangent_bundle_core I M).to_vector_bundle_core).is_open_map_proj - -/-- In the tangent bundle to the model space, the charts are just the canonical identification -between a product type and a sigma type, a.k.a. `equiv.sigma_equiv_prod`. -/ -@[simp, mfld_simps] lemma tangent_bundle_model_space_chart_at (p : tangent_bundle I H) : - (chart_at (model_prod H E) p).to_local_equiv = (equiv.sigma_equiv_prod H E).to_local_equiv := -begin - have A : ∀ x_fst, fderiv_within 𝕜 (I ∘ I.symm) (range I) (I x_fst) = continuous_linear_map.id 𝕜 E, - { assume x_fst, - have : fderiv_within 𝕜 (I ∘ I.symm) (range I) (I x_fst) - = fderiv_within 𝕜 id (range I) (I x_fst), - { refine fderiv_within_congr I.unique_diff_at_image (λ y hy, _) (by simp), - exact model_with_corners.right_inv _ hy }, - rwa fderiv_within_id I.unique_diff_at_image at this }, - ext x : 1, - show (chart_at (model_prod H E) p : tangent_bundle I H → model_prod H E) x = - (equiv.sigma_equiv_prod H E) x, - { cases x, - simp only [chart_at, basic_smooth_vector_bundle_core.chart, tangent_bundle_core, - basic_smooth_vector_bundle_core.to_vector_bundle_core, A, prod.mk.inj_iff, - continuous_linear_map.coe_id'] with mfld_simps }, - show ∀ x, ((chart_at (model_prod H E) p).to_local_equiv).symm x = - (equiv.sigma_equiv_prod H E).symm x, - { rintros ⟨x_fst, x_snd⟩, - simp only [basic_smooth_vector_bundle_core.to_vector_bundle_core, - tangent_bundle_core, A, continuous_linear_map.coe_id', basic_smooth_vector_bundle_core.chart, - chart_at, continuous_linear_map.coe_coe, sigma.mk.inj_iff] with mfld_simps, }, - show ((chart_at (model_prod H E) p).to_local_equiv).source = univ, - by simp only [chart_at] with mfld_simps, -end - -@[simp, mfld_simps] lemma tangent_bundle_model_space_coe_chart_at (p : tangent_bundle I H) : - ⇑(chart_at (model_prod H E) p) = equiv.sigma_equiv_prod H E := -by { unfold_coes, simp only with mfld_simps } - -@[simp, mfld_simps] lemma tangent_bundle_model_space_coe_chart_at_symm (p : tangent_bundle I H) : - ((chart_at (model_prod H E) p).symm : model_prod H E → tangent_bundle I H) = - (equiv.sigma_equiv_prod H E).symm := -by { unfold_coes, simp only with mfld_simps } - -variable (H) -/-- The canonical identification between the tangent bundle to the model space and the product, -as a homeomorphism -/ -def tangent_bundle_model_space_homeomorph : tangent_bundle I H ≃ₜ model_prod H E := -{ continuous_to_fun := - begin - let p : tangent_bundle I H := ⟨I.symm (0 : E), (0 : E)⟩, - have : continuous (chart_at (model_prod H E) p), - { rw continuous_iff_continuous_on_univ, - convert local_homeomorph.continuous_on _, - simp only with mfld_simps }, - simpa only with mfld_simps using this, - end, - continuous_inv_fun := - begin - let p : tangent_bundle I H := ⟨I.symm (0 : E), (0 : E)⟩, - have : continuous (chart_at (model_prod H E) p).symm, - { rw continuous_iff_continuous_on_univ, - convert local_homeomorph.continuous_on _, - simp only with mfld_simps }, - simpa only with mfld_simps using this, - end, - .. equiv.sigma_equiv_prod H E } - -@[simp, mfld_simps] lemma tangent_bundle_model_space_homeomorph_coe : - (tangent_bundle_model_space_homeomorph H I : tangent_bundle I H → model_prod H E) - = equiv.sigma_equiv_prod H E := -rfl - -@[simp, mfld_simps] lemma tangent_bundle_model_space_homeomorph_coe_symm : - ((tangent_bundle_model_space_homeomorph H I).symm : model_prod H E → tangent_bundle I H) - = (equiv.sigma_equiv_prod H E).symm := -rfl - -end tangent_bundle diff --git a/src/geometry/manifold/vector_bundle/basic.lean b/src/geometry/manifold/vector_bundle/basic.lean index 592e326b194a9..19ea2d529b202 100644 --- a/src/geometry/manifold/vector_bundle/basic.lean +++ b/src/geometry/manifold/vector_bundle/basic.lean @@ -54,8 +54,8 @@ fields, they can also be C^k vector bundles, etc. assert_not_exists mfderiv -open bundle set local_homeomorph -open_locale manifold bundle +open bundle set local_homeomorph function (id_def) filter +open_locale manifold bundle topology variables {𝕜 B B' F M : Type*} {E : B → Type*} @@ -95,6 +95,144 @@ begin trivialization.coe_fst' _ (mem_base_set_trivialization_at F E x.proj)] end +lemma fiber_bundle.charted_space_chart_at_symm_fst (x : total_space E) (y : model_prod HB F) + (hy : y ∈ (chart_at (model_prod HB F) x).target) : + ((chart_at (model_prod HB F) x).symm y).proj = (chart_at HB x.proj).symm y.1 := +begin + simp only [fiber_bundle.charted_space_chart_at] with mfld_simps at hy ⊢, + exact (trivialization_at F E x.proj).proj_symm_apply hy.2, +end + +end + +section +variables [nontrivially_normed_field 𝕜] + [normed_add_comm_group F] [normed_space 𝕜 F] + [topological_space (total_space E)] [∀ x, topological_space (E x)] + + {EB : Type*} [normed_add_comm_group EB] [normed_space 𝕜 EB] + {HB : Type*} [topological_space HB] (IB : model_with_corners 𝕜 EB HB) + (E' : B → Type*) [Π x, has_zero (E' x)] + {EM : Type*} [normed_add_comm_group EM] [normed_space 𝕜 EM] + {HM : Type*} [topological_space HM] {IM : model_with_corners 𝕜 EM HM} + [topological_space M] [charted_space HM M] [Is : smooth_manifold_with_corners IM M] + {n : ℕ∞} + +variables [topological_space B] [charted_space HB B] [fiber_bundle F E] + +protected lemma fiber_bundle.ext_chart_at (x : total_space E) : + ext_chart_at (IB.prod 𝓘(𝕜, F)) x = + (trivialization_at F E x.proj).to_local_equiv ≫ + (ext_chart_at IB x.proj).prod (local_equiv.refl F) := +begin + simp_rw [ext_chart_at, fiber_bundle.charted_space_chart_at, extend], + simp only [local_equiv.trans_assoc] with mfld_simps, +end + +/-! ### Smoothness of maps in/out fiber bundles + +Note: For these results we don't need that the bundle is a smooth vector bundle, or even a vector +bundle at all, just that it is a fiber bundle over a charted base space. +-/ + +namespace bundle +variables {F E IB} + +/-- Characterization of C^n functions into a smooth vector bundle. -/ +lemma cont_mdiff_within_at_total_space (f : M → total_space E) {s : set M} {x₀ : M} : + cont_mdiff_within_at IM (IB.prod (𝓘(𝕜, F))) n f s x₀ ↔ + cont_mdiff_within_at IM IB n (λ x, (f x).proj) s x₀ ∧ + cont_mdiff_within_at IM 𝓘(𝕜, F) n (λ x, (trivialization_at F E (f x₀).proj (f x)).2) s x₀ := +begin + simp only [cont_mdiff_within_at_iff_target] {single_pass := tt}, + rw [and_and_and_comm, ← continuous_within_at_total_space, and.congr_right_iff], + intros hf, + simp_rw [model_with_corners_self_prod, fiber_bundle.ext_chart_at, function.comp, + local_equiv.trans_apply, local_equiv.prod_coe, local_equiv.refl_coe, + ext_chart_at_self_apply, model_with_corners_self_coe, id_def], + refine (cont_mdiff_within_at_prod_iff _).trans _, -- rw doesn't do this? + have h1 : (λ x, (f x).proj) ⁻¹' (trivialization_at F E (f x₀).proj).base_set ∈ 𝓝[s] x₀ := + ((continuous_proj F E).continuous_within_at.comp hf (maps_to_image f s)) + .preimage_mem_nhds_within + ((trivialization.open_base_set _).mem_nhds (mem_base_set_trivialization_at F E _)), + refine and_congr (eventually_eq.cont_mdiff_within_at_iff (eventually_of_mem h1 $ λ x hx, _) _) + iff.rfl, + { simp_rw [function.comp, local_homeomorph.coe_coe, trivialization.coe_coe], + rw [trivialization.coe_fst'], + exact hx }, + { simp only with mfld_simps }, +end + +/-- Characterization of C^n functions into a smooth vector bundle. -/ +lemma cont_mdiff_at_total_space (f : M → total_space E) (x₀ : M) : + cont_mdiff_at IM (IB.prod (𝓘(𝕜, F))) n f x₀ ↔ + cont_mdiff_at IM IB n (λ x, (f x).proj) x₀ ∧ + cont_mdiff_at IM 𝓘(𝕜, F) n (λ x, (trivialization_at F E (f x₀).proj (f x)).2) x₀ := +by { simp_rw [← cont_mdiff_within_at_univ], exact cont_mdiff_within_at_total_space f } + +variables (E) +lemma cont_mdiff_proj : cont_mdiff (IB.prod 𝓘(𝕜, F)) IB n (π E) := +begin + intro x, + rw [cont_mdiff_at, cont_mdiff_within_at_iff'], + refine ⟨(continuous_proj F E).continuous_within_at, _⟩, + simp_rw [(∘), fiber_bundle.ext_chart_at], + apply cont_diff_within_at_fst.congr, + { rintros ⟨a, b⟩ hab, + simp only with mfld_simps at hab, + have : ((chart_at HB x.1).symm (IB.symm a), b) ∈ (trivialization_at F E x.fst).target, + { simp only [hab] with mfld_simps }, + simp only [trivialization.proj_symm_apply _ this, hab] with mfld_simps }, + { simp only with mfld_simps } +end + +lemma smooth_proj : smooth (IB.prod 𝓘(𝕜, F)) IB (π E) := +cont_mdiff_proj E + +lemma cont_mdiff_on_proj {s : set (total_space E)} : + cont_mdiff_on (IB.prod 𝓘(𝕜, F)) IB n (π E) s := +(bundle.cont_mdiff_proj E).cont_mdiff_on + +lemma smooth_on_proj {s : set (total_space E)} : + smooth_on (IB.prod 𝓘(𝕜, F)) IB (π E) s := +cont_mdiff_on_proj E + +lemma cont_mdiff_at_proj {p : total_space E} : + cont_mdiff_at (IB.prod 𝓘(𝕜, F)) IB n + (π E) p := +(bundle.cont_mdiff_proj E).cont_mdiff_at + +lemma smooth_at_proj {p : total_space E} : + smooth_at (IB.prod 𝓘(𝕜, F)) IB (π E) p := +bundle.cont_mdiff_at_proj E + +lemma cont_mdiff_within_at_proj + {s : set (total_space E)} + {p : total_space E} : + cont_mdiff_within_at (IB.prod 𝓘(𝕜, F)) IB n (π E) s p := +(bundle.cont_mdiff_at_proj E).cont_mdiff_within_at + +lemma smooth_within_at_proj + {s : set (total_space E)} + {p : total_space E} : + smooth_within_at (IB.prod 𝓘(𝕜, F)) IB (π E) s p := +bundle.cont_mdiff_within_at_proj E + +variables (𝕜 E) [∀ x, add_comm_monoid (E x)] [∀ x, module 𝕜 (E x)] [vector_bundle 𝕜 F E] + +lemma smooth_zero_section : smooth IB (IB.prod 𝓘(𝕜, F)) (zero_section E) := +begin + intro x, + rw [bundle.cont_mdiff_at_total_space], + refine ⟨cont_mdiff_at_id, cont_mdiff_at_const.congr_of_eventually_eq _⟩, + { exact 0 }, + refine eventually_of_mem ((trivialization_at F E x).open_base_set.mem_nhds + (mem_base_set_trivialization_at F E x)) (λ x' hx', _), + simp_rw [zero_section_proj, (trivialization_at F E x).zero_section 𝕜 hx'] +end + +end bundle + end /-! ### Smooth vector bundles -/ @@ -106,6 +244,10 @@ variables [nontrivially_normed_field 𝕜] [∀ x, add_comm_monoid (E x)] [∀ x {EB : Type*} [normed_add_comm_group EB] [normed_space 𝕜 EB] {HB : Type*} [topological_space HB] (IB : model_with_corners 𝕜 EB HB) [topological_space B] [charted_space HB B] [smooth_manifold_with_corners IB B] + {EM : Type*} [normed_add_comm_group EM] [normed_space 𝕜 EM] + {HM : Type*} [topological_space HM] {IM : model_with_corners 𝕜 EM HM} + [topological_space M] [charted_space HM M] [Is : smooth_manifold_with_corners IM M] + {n : ℕ∞} variables (F E) [fiber_bundle F E] [vector_bundle 𝕜 F E] @@ -123,7 +265,6 @@ export smooth_vector_bundle (smooth_on_coord_change) variables [smooth_vector_bundle F E IB] - /-- For a smooth vector bundle `E` over `B` with fiber modelled on `F`, the change-of-co-ordinates between two trivializations `e`, `e'` for `E`, considered as charts to `B × F`, is smooth and fiberwise linear. -/ diff --git a/src/geometry/manifold/vector_bundle/tangent.lean b/src/geometry/manifold/vector_bundle/tangent.lean index 08c238a379f5d..870a40386d099 100644 --- a/src/geometry/manifold/vector_bundle/tangent.lean +++ b/src/geometry/manifold/vector_bundle/tangent.lean @@ -30,7 +30,7 @@ This defines a smooth vector bundle `tangent_bundle` with fibers `tangent_space` -/ open bundle set smooth_manifold_with_corners local_homeomorph -open_locale manifold topology +open_locale manifold topology bundle noncomputable theory @@ -41,9 +41,6 @@ variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] {F : Type*} [normed_add_comm_group F] [normed_space 𝕜 F] variables (I) -namespace hidden /- we temporarily put everything in a namespace to not have name clashes with -the existing `tangent_bundle_core`. -/ - /-- Auxiliary lemma for tangent spaces: the derivative of a coordinate change between two charts is smooth on its source. -/ lemma cont_diff_on_fderiv_coord_change (i j : atlas H M) : @@ -154,26 +151,76 @@ instance : inhabited (tangent_space I x) := ⟨0⟩ end instance : topological_space TM := -(tangent_bundle_core I M).to_fiber_bundle_core.to_topological_space +(tangent_bundle_core I M).to_topological_space instance : fiber_bundle E (tangent_space I : M → Type*) := -(tangent_bundle_core I M).to_fiber_bundle_core.fiber_bundle +(tangent_bundle_core I M).fiber_bundle instance : vector_bundle 𝕜 E (tangent_space I : M → Type*) := (tangent_bundle_core I M).vector_bundle -lemma tangent_space_chart_at (p : TM) : +namespace tangent_bundle + +protected lemma chart_at (p : TM) : chart_at (model_prod H E) p = ((tangent_bundle_core I M).to_fiber_bundle_core.local_triv (achart H p.1)) .to_local_homeomorph ≫ₕ (chart_at H p.1).prod (local_homeomorph.refl E) := rfl -lemma tangent_space_chart_at_to_local_equiv (p : TM) : +lemma chart_at_to_local_equiv (p : TM) : (chart_at (model_prod H E) p).to_local_equiv = (tangent_bundle_core I M).to_fiber_bundle_core.local_triv_as_local_equiv (achart H p.1) ≫ (chart_at H p.1).to_local_equiv.prod (local_equiv.refl E) := rfl +lemma trivialization_at_eq_local_triv (x : M) : + trivialization_at E (tangent_space I) x = + (tangent_bundle_core I M).to_fiber_bundle_core.local_triv (achart H x) := +rfl + +@[simp, mfld_simps] +lemma trivialization_at_source (x : M) : + (trivialization_at E (tangent_space I) x).source = π _ ⁻¹' (chart_at H x).source := +rfl + +@[simp, mfld_simps] +lemma trivialization_at_target (x : M) : + (trivialization_at E (tangent_space I) x).target = (chart_at H x).source ×ˢ univ := +rfl + +@[simp, mfld_simps] +lemma trivialization_at_base_set (x : M) : + (trivialization_at E (tangent_space I) x).base_set = (chart_at H x).source := +rfl + +lemma trivialization_at_apply (x : M) (z : TM) : + trivialization_at E (tangent_space I) x z = + (z.1, fderiv_within 𝕜 ((chart_at H x).extend I ∘ ((chart_at H z.1).extend I).symm) (range I) + ((chart_at H z.1).extend I z.1) z.2) := +rfl + +@[simp, mfld_simps] +lemma trivialization_at_fst (x : M) (z : TM) : + (trivialization_at E (tangent_space I) x z).1 = z.1 := +rfl + +@[simp, mfld_simps] lemma mem_chart_source_iff (p q : TM) : + p ∈ (chart_at (model_prod H E) q).source ↔ p.1 ∈ (chart_at H q.1).source := +by simp only [fiber_bundle.charted_space_chart_at] with mfld_simps + +@[simp, mfld_simps] lemma mem_chart_target_iff (p : H × E) (q : TM) : + p ∈ (chart_at (model_prod H E) q).target ↔ p.1 ∈ (chart_at H q.1).target := +by simp only [fiber_bundle.charted_space_chart_at, and_iff_left_iff_imp] with mfld_simps + {contextual := tt} + +@[simp, mfld_simps] lemma coe_chart_at_fst (p q : TM) : + ((chart_at (model_prod H E) q) p).1 = chart_at H q.1 p.1 := rfl + +@[simp, mfld_simps] lemma coe_chart_at_symm_fst (p : H × E) (q : TM) : + ((chart_at (model_prod H E) q).symm p).1 = ((chart_at H q.1).symm : H → M) p.1 := rfl + +end tangent_bundle + instance tangent_bundle_core.is_smooth : (tangent_bundle_core I M).is_smooth I := begin refine ⟨λ i j, _⟩, @@ -207,7 +254,7 @@ begin { intros x, ext, { refl }, apply heq_of_eq, exact (tangent_bundle_core I H).coord_change_self (achart _ x.1) x.1 (mem_achart_source H x.1) x.2 }, - simp_rw [tangent_space_chart_at, fiber_bundle_core.local_triv, + simp_rw [tangent_bundle.chart_at, fiber_bundle_core.local_triv, fiber_bundle_core.local_triv_as_local_equiv, vector_bundle_core.to_fiber_bundle_core_base_set, tangent_bundle_core_base_set], simp only with mfld_simps, @@ -256,5 +303,3 @@ rfl ((tangent_bundle_model_space_homeomorph H I).symm : model_prod H E → tangent_bundle I H) = (equiv.sigma_equiv_prod H E).symm := rfl - -end hidden diff --git a/src/topology/fiber_bundle/basic.lean b/src/topology/fiber_bundle/basic.lean index 48ad8e2aba735..f9271f8899861 100644 --- a/src/topology/fiber_bundle/basic.lean +++ b/src/topology/fiber_bundle/basic.lean @@ -167,12 +167,12 @@ for the initial bundle. Fiber bundle, topological bundle, structure group -/ -variables {ι : Type*} {B : Type*} {F : Type*} +variables {ι B F X : Type*} [topological_space X] open topological_space filter set bundle open_locale topology classical bundle -attribute [mfld_simps] total_space.proj total_space_mk coe_fst coe_snd coe_snd_map_apply +attribute [mfld_simps] total_space_mk coe_fst coe_snd coe_snd_map_apply coe_snd_map_smul total_space.mk_cast /-! ### General definition of fiber bundles -/ @@ -240,6 +240,47 @@ lemma quotient_map_proj [nonempty F] : quotient_map (π E) := lemma continuous_total_space_mk (x : B) : continuous (@total_space_mk B E x) := (total_space_mk_inducing F E x).continuous +variables {E F} + +@[simp, mfld_simps] +lemma mem_trivialization_at_proj_source {x : total_space E} : + x ∈ (trivialization_at F E x.proj).source := +(trivialization.mem_source _).mpr $ mem_base_set_trivialization_at F E x.proj + +@[simp, mfld_simps] +lemma trivialization_at_proj_fst {x : total_space E} : + ((trivialization_at F E x.proj) x).1 = x.proj := +trivialization.coe_fst' _ $ mem_base_set_trivialization_at F E x.proj + +variable (F) +open trivialization + +/-- Characterization of continuous functions (at a point, within a set) into a fiber bundle. -/ +lemma continuous_within_at_total_space (f : X → total_space E) {s : set X} {x₀ : X} : + continuous_within_at f s x₀ ↔ + continuous_within_at (λ x, (f x).proj) s x₀ ∧ + continuous_within_at (λ x, ((trivialization_at F E (f x₀).proj) (f x)).2) s x₀ := +begin + refine (and_iff_right_iff_imp.2 $ λ hf, _).symm.trans (and_congr_right $ λ hf, _), + { refine (continuous_proj F E).continuous_within_at.comp hf (maps_to_image f s) }, + have h1 : (λ x, (f x).proj) ⁻¹' (trivialization_at F E (f x₀).proj).base_set ∈ 𝓝[s] x₀ := + hf.preimage_mem_nhds_within ((open_base_set _).mem_nhds (mem_base_set_trivialization_at F E _)), + have h2 : continuous_within_at (λ x, (trivialization_at F E (f x₀).proj (f x)).1) s x₀, + { refine hf.congr_of_eventually_eq (eventually_of_mem h1 $ λ x hx, _) trivialization_at_proj_fst, + rw [coe_fst'], + exact hx }, + rw [(trivialization_at F E (f x₀).proj).continuous_within_at_iff_continuous_within_at_comp_left], + { simp_rw [continuous_within_at_prod_iff, function.comp, trivialization.coe_coe, h2, true_and] }, + { apply mem_trivialization_at_proj_source }, + { rwa [source_eq, preimage_preimage] } +end + +/-- Characterization of continuous functions (at a point) into a fiber bundle. -/ +lemma continuous_at_total_space (f : X → total_space E) {x₀ : X} : + continuous_at f x₀ ↔ continuous_at (λ x, (f x).proj) x₀ ∧ + continuous_at (λ x, ((trivialization_at F E (f x₀).proj) (f x)).2) x₀ := +by { simp_rw [← continuous_within_at_univ], exact continuous_within_at_total_space F f } + end fiber_bundle variables (F E) diff --git a/src/topology/vector_bundle/basic.lean b/src/topology/vector_bundle/basic.lean index 2efc8c158d5e9..235477250dbab 100644 --- a/src/topology/vector_bundle/basic.lean +++ b/src/topology/vector_bundle/basic.lean @@ -320,6 +320,20 @@ end topological_vector_space section +namespace bundle + +/-- The zero section of a vector bundle -/ +def zero_section [∀ x, has_zero (E x)] : B → total_space E := +λ x, total_space_mk x 0 + +@[simp, mfld_simps] +lemma zero_section_proj [∀ x, has_zero (E x)] (x : B) : (zero_section E x).proj = x := rfl +@[simp, mfld_simps] +lemma zero_section_snd [∀ x, has_zero (E x)] (x : B) : (zero_section E x).2 = 0 := rfl + +end bundle +open bundle + variables [nontrivially_normed_field R] [∀ x, add_comm_monoid (E x)] [∀ x, module R (E x)] [normed_add_comm_group F] [normed_space R F] [topological_space B] [topological_space (total_space E)] [∀ x, topological_space (E x)] [fiber_bundle F E] @@ -363,7 +377,7 @@ def continuous_linear_map_at (e : trivialization F (π E)) [e.is_linear R] (b : dsimp, rw [e.coe_linear_map_at b], refine continuous_if_const _ (λ hb, _) (λ _, continuous_zero), - exact continuous_snd.comp (e.to_local_homeomorph.continuous_on.comp_continuous + exact continuous_snd.comp (e.continuous_on.comp_continuous (fiber_bundle.total_space_mk_inducing F E b).continuous (λ x, e.mem_source.mpr hb)) end, @@ -403,7 +417,7 @@ def continuous_linear_equiv_at (e : trivialization F (π E)) [e.is_linear R] (b (hb : b ∈ e.base_set) : E b ≃L[R] F := { to_fun := λ y, (e (total_space_mk b y)).2, -- given explicitly to help `simps` inv_fun := e.symm b, -- given explicitly to help `simps` - continuous_to_fun := continuous_snd.comp (e.to_local_homeomorph.continuous_on.comp_continuous + continuous_to_fun := continuous_snd.comp (e.continuous_on.comp_continuous (fiber_bundle.total_space_mk_inducing F E b).continuous (λ x, e.mem_source.mpr hb)), continuous_inv_fun := (e.symmL R b).continuous, @@ -429,7 +443,7 @@ variables (R) lemma apply_eq_prod_continuous_linear_equiv_at (e : trivialization F (π E)) [e.is_linear R] (b : B) (hb : b ∈ e.base_set) (z : E b) : - e.to_local_homeomorph ⟨b, z⟩ = (b, e.continuous_linear_equiv_at R b hb z) := + e ⟨b, z⟩ = (b, e.continuous_linear_equiv_at R b hb z) := begin ext, { refine e.coe_fst _, @@ -438,6 +452,11 @@ begin { simp only [coe_coe, continuous_linear_equiv_at_apply] } end +protected lemma zero_section (e : trivialization F (π E)) [e.is_linear R] + {x : B} (hx : x ∈ e.base_set) : e (zero_section E x) = (x, 0) := +by simp_rw [zero_section, total_space_mk, e.apply_eq_prod_continuous_linear_equiv_at R x hx 0, + map_zero] + variables {R} lemma symm_apply_eq_mk_continuous_linear_equiv_at_symm (e : trivialization F (π E)) [e.is_linear R] @@ -445,12 +464,12 @@ lemma symm_apply_eq_mk_continuous_linear_equiv_at_symm (e : trivialization F (π e.to_local_homeomorph.symm ⟨b, z⟩ = total_space_mk b ((e.continuous_linear_equiv_at R b hb).symm z) := begin - have h : (b, z) ∈ e.to_local_homeomorph.target, + have h : (b, z) ∈ e.target, { rw e.target_eq, exact ⟨hb, mem_univ _⟩ }, - apply e.to_local_homeomorph.inj_on (e.to_local_homeomorph.map_target h), - { simp only [e.source_eq, hb, mem_preimage]}, - simp_rw [e.apply_eq_prod_continuous_linear_equiv_at R b hb, e.to_local_homeomorph.right_inv h, + apply e.inj_on (e.map_target h), + { simp only [e.source_eq, hb, mem_preimage] }, + simp_rw [e.right_inv h, coe_coe, e.apply_eq_prod_continuous_linear_equiv_at R b hb, continuous_linear_equiv.apply_symm_apply], end From e8ac6315bcfcbaf2d19a046719c3b553206dac75 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 28 Mar 2023 23:45:16 +0000 Subject: [PATCH 0712/1291] chore(order/category): Rename categories (#18657) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The Lean 4 naming convention forces us to rename the categories of orders. Instead of blankly appending `Cat` to the names, we proactively shorten the names. Incidentally, this gets them closer to the way they're referred in the literature. * `Preorder` → `Preord` (the literature name is `Ord`, but Lean 4 already takes it) * `PartialOrder` → `PartOrd` * `BoundedOrder` → `BddOrd` * `FinPartialOrder` → `FinPartOrd` * `SemilatticeSup` → `SemilatSup` * `SemilatticeInf` → `SemilatInf` * `Lattice` → `Lat` * `DistribLattice` → `DistLat` * `BoundedLattice` → `BddLat` * `BoundedDistribLattice` → `BddDistLat` * `LinearOrder` → `LinOrd` * `CompleteLattice` → `CompleteLat` * `Frame` → `Frm` (the corresponding class is `Order.Frame`, but better be safe) Co-authored-by: Jeremy Tan Jie Rui --- src/algebraic_topology/simplex_category.lean | 5 +- src/order/category/BddDistLat.lean | 83 ++++++++++ src/order/category/BddLat.lean | 125 +++++++++++++++ src/order/category/BddOrd.lean | 83 ++++++++++ src/order/category/BoolAlg.lean | 20 +-- src/order/category/BoundedDistribLattice.lean | 83 ---------- src/order/category/BoundedLattice.lean | 125 --------------- src/order/category/BoundedOrder.lean | 83 ---------- src/order/category/CompleteLat.lean | 68 ++++++++ src/order/category/CompleteLattice.lean | 68 -------- src/order/category/DistLat.lean | 65 ++++++++ src/order/category/DistribLattice.lean | 65 -------- src/order/category/FinBoolAlg.lean | 8 +- src/order/category/FinPartOrd.lean | 79 ++++++++++ src/order/category/FinPartialOrder.lean | 79 ---------- src/order/category/{Frame.lean => Frm.lean} | 32 ++-- src/order/category/HeytAlg.lean | 6 +- src/order/category/Lat.lean | 75 +++++++++ src/order/category/Lattice.lean | 75 --------- src/order/category/LinOrd.lean | 64 ++++++++ src/order/category/LinearOrder.lean | 64 -------- src/order/category/NonemptyFinLinOrd.lean | 12 +- .../{PartialOrder.lean => PartOrd.lean} | 60 +++---- .../category/{Preorder.lean => Preord.lean} | 36 ++--- src/order/category/Semilat.lean | 148 ++++++++++++++++++ src/order/category/Semilattice.lean | 148 ------------------ src/order/complete_boolean_algebra.lean | 2 +- src/order/hom/complete_lattice.lean | 6 +- src/topology/category/Locale.lean | 6 +- 29 files changed, 886 insertions(+), 887 deletions(-) create mode 100644 src/order/category/BddDistLat.lean create mode 100644 src/order/category/BddLat.lean create mode 100644 src/order/category/BddOrd.lean delete mode 100644 src/order/category/BoundedDistribLattice.lean delete mode 100644 src/order/category/BoundedLattice.lean delete mode 100644 src/order/category/BoundedOrder.lean create mode 100644 src/order/category/CompleteLat.lean delete mode 100644 src/order/category/CompleteLattice.lean create mode 100644 src/order/category/DistLat.lean delete mode 100644 src/order/category/DistribLattice.lean create mode 100644 src/order/category/FinPartOrd.lean delete mode 100644 src/order/category/FinPartialOrder.lean rename src/order/category/{Frame.lean => Frm.lean} (69%) create mode 100644 src/order/category/Lat.lean delete mode 100644 src/order/category/Lattice.lean create mode 100644 src/order/category/LinOrd.lean delete mode 100644 src/order/category/LinearOrder.lean rename src/order/category/{PartialOrder.lean => PartOrd.lean} (52%) rename src/order/category/{Preorder.lean => Preord.lean} (62%) create mode 100644 src/order/category/Semilat.lean delete mode 100644 src/order/category/Semilattice.lean diff --git a/src/algebraic_topology/simplex_category.lean b/src/algebraic_topology/simplex_category.lean index e8c9922574dc3..f331834ebe631 100644 --- a/src/algebraic_topology/simplex_category.lean +++ b/src/algebraic_topology/simplex_category.lean @@ -811,8 +811,7 @@ end epi_mono to the category attached to the ordered set `{0, 1, ..., n}` -/ @[simps obj map] def to_Cat : simplex_category ⥤ Cat.{0} := -simplex_category.skeletal_functor ⋙ forget₂ NonemptyFinLinOrd LinearOrder ⋙ - forget₂ LinearOrder Lattice ⋙ forget₂ Lattice PartialOrder ⋙ - forget₂ PartialOrder Preorder ⋙ Preorder_to_Cat +simplex_category.skeletal_functor ⋙ forget₂ NonemptyFinLinOrd LinOrd ⋙ + forget₂ LinOrd Lat ⋙ forget₂ Lat PartOrd ⋙ forget₂ PartOrd Preord ⋙ Preord_to_Cat end simplex_category diff --git a/src/order/category/BddDistLat.lean b/src/order/category/BddDistLat.lean new file mode 100644 index 0000000000000..7523cd5d45659 --- /dev/null +++ b/src/order/category/BddDistLat.lean @@ -0,0 +1,83 @@ +/- +Copyright (c) 2022 Yaël Dillies. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies +-/ +import order.category.BddLat +import order.category.DistLat + +/-! +# The category of bounded distributive lattices + +This defines `BddDistLat`, the category of bounded distributive lattices. + +Note that this category is sometimes called [`DistLat`](https://ncatlab.org/nlab/show/DistLat) when +being a lattice is understood to entail having a bottom and a top element. +-/ + +universes u + +open category_theory + +/-- The category of bounded distributive lattices with bounded lattice morphisms. -/ +structure BddDistLat := +(to_DistLat : DistLat) +[is_bounded_order : bounded_order to_DistLat] + +namespace BddDistLat + +instance : has_coe_to_sort BddDistLat Type* := ⟨λ X, X.to_DistLat⟩ +instance (X : BddDistLat) : distrib_lattice X := X.to_DistLat.str + +attribute [instance] BddDistLat.is_bounded_order + +/-- Construct a bundled `BddDistLat` from a `bounded_order` `distrib_lattice`. -/ +def of (α : Type*) [distrib_lattice α] [bounded_order α] : BddDistLat := ⟨⟨α⟩⟩ + +@[simp] lemma coe_of (α : Type*) [distrib_lattice α] [bounded_order α] : ↥(of α) = α := rfl + +instance : inhabited BddDistLat := ⟨of punit⟩ + +/-- Turn a `BddDistLat` into a `BddLat` by forgetting it is distributive. -/ +def to_BddLat (X : BddDistLat) : BddLat := BddLat.of X + +@[simp] lemma coe_to_BddLat (X : BddDistLat) : ↥X.to_BddLat = ↥X := rfl + +instance : large_category.{u} BddDistLat := induced_category.category to_BddLat + +instance : concrete_category BddDistLat := +induced_category.concrete_category to_BddLat + +instance has_forget_to_DistLat : has_forget₂ BddDistLat DistLat := +{ forget₂ := { obj := λ X, ⟨X⟩, map := λ X Y, bounded_lattice_hom.to_lattice_hom } } + +instance has_forget_to_BddLat : has_forget₂ BddDistLat BddLat := +induced_category.has_forget₂ to_BddLat + +lemma forget_BddLat_Lat_eq_forget_DistLat_Lat : + forget₂ BddDistLat BddLat ⋙ forget₂ BddLat Lat = + forget₂ BddDistLat DistLat ⋙ forget₂ DistLat Lat := rfl + +/-- Constructs an equivalence between bounded distributive lattices from an order isomorphism +between them. -/ +@[simps] def iso.mk {α β : BddDistLat.{u}} (e : α ≃o β) : α ≅ β := +{ hom := (e : bounded_lattice_hom α β), + inv := (e.symm : bounded_lattice_hom β α), + hom_inv_id' := by { ext, exact e.symm_apply_apply _ }, + inv_hom_id' := by { ext, exact e.apply_symm_apply _ } } + +/-- `order_dual` as a functor. -/ +@[simps] def dual : BddDistLat ⥤ BddDistLat := +{ obj := λ X, of Xᵒᵈ, map := λ X Y, bounded_lattice_hom.dual } + +/-- The equivalence between `BddDistLat` and itself induced by `order_dual` both ways. -/ +@[simps functor inverse] def dual_equiv : BddDistLat ≌ BddDistLat := +equivalence.mk dual dual + (nat_iso.of_components (λ X, iso.mk $ order_iso.dual_dual X) $ λ X Y f, rfl) + (nat_iso.of_components (λ X, iso.mk $ order_iso.dual_dual X) $ λ X Y f, rfl) + +end BddDistLat + +lemma BddDistLat_dual_comp_forget_to_DistLat : + BddDistLat.dual ⋙ forget₂ BddDistLat DistLat = + forget₂ BddDistLat DistLat ⋙ DistLat.dual := rfl diff --git a/src/order/category/BddLat.lean b/src/order/category/BddLat.lean new file mode 100644 index 0000000000000..f6166c180a2a9 --- /dev/null +++ b/src/order/category/BddLat.lean @@ -0,0 +1,125 @@ +/- +Copyright (c) 2022 Yaël Dillies. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies +-/ +import order.category.BddOrd +import order.category.Lat +import order.category.Semilat + +/-! +# The category of bounded lattices + +This file defines `BddLat`, the category of bounded lattices. + +In literature, this is sometimes called `Lat`, the category of lattices, because being a lattice is +understood to entail having a bottom and a top element. +-/ + +universes u + +open category_theory + +/-- The category of bounded lattices with bounded lattice morphisms. -/ +structure BddLat := +(to_Lat : Lat) +[is_bounded_order : bounded_order to_Lat] + +namespace BddLat + +instance : has_coe_to_sort BddLat Type* := ⟨λ X, X.to_Lat⟩ +instance (X : BddLat) : lattice X := X.to_Lat.str + +attribute [instance] BddLat.is_bounded_order + +/-- Construct a bundled `BddLat` from `lattice` + `bounded_order`. -/ +def of (α : Type*) [lattice α] [bounded_order α] : BddLat := ⟨⟨α⟩⟩ + +@[simp] lemma coe_of (α : Type*) [lattice α] [bounded_order α] : ↥(of α) = α := rfl + +instance : inhabited BddLat := ⟨of punit⟩ + +instance : large_category.{u} BddLat := +{ hom := λ X Y, bounded_lattice_hom X Y, + id := λ X, bounded_lattice_hom.id X, + comp := λ X Y Z f g, g.comp f, + id_comp' := λ X Y, bounded_lattice_hom.comp_id, + comp_id' := λ X Y, bounded_lattice_hom.id_comp, + assoc' := λ W X Y Z _ _ _, bounded_lattice_hom.comp_assoc _ _ _ } + +instance : concrete_category BddLat := +{ forget := ⟨coe_sort, λ X Y, coe_fn, λ X, rfl, λ X Y Z f g, rfl⟩, + forget_faithful := ⟨λ X Y, by convert fun_like.coe_injective⟩ } + +instance has_forget_to_BddOrd : has_forget₂ BddLat BddOrd := +{ forget₂ := { obj := λ X, BddOrd.of X, + map := λ X Y, bounded_lattice_hom.to_bounded_order_hom } } + +instance has_forget_to_Lat : has_forget₂ BddLat Lat := +{ forget₂ := { obj := λ X, ⟨X⟩, map := λ X Y, bounded_lattice_hom.to_lattice_hom } } + +instance has_forget_to_SemilatSup : has_forget₂ BddLat SemilatSup := +{ forget₂ := { obj := λ X, ⟨X⟩, map := λ X Y, bounded_lattice_hom.to_sup_bot_hom } } + +instance has_forget_to_SemilatInf : has_forget₂ BddLat SemilatInf := +{ forget₂ := { obj := λ X, ⟨X⟩, map := λ X Y, bounded_lattice_hom.to_inf_top_hom } } + +@[simp] lemma coe_forget_to_BddOrd (X : BddLat) : + ↥((forget₂ BddLat BddOrd).obj X) = ↥X := rfl + +@[simp] lemma coe_forget_to_Lat (X : BddLat) : + ↥((forget₂ BddLat Lat).obj X) = ↥X := rfl + +@[simp] lemma coe_forget_to_SemilatSup (X : BddLat) : + ↥((forget₂ BddLat SemilatSup).obj X) = ↥X := rfl + +@[simp] lemma coe_forget_to_SemilatInf (X : BddLat) : + ↥((forget₂ BddLat SemilatInf).obj X) = ↥X := rfl + +lemma forget_Lat_PartOrd_eq_forget_BddOrd_PartOrd : + forget₂ BddLat Lat ⋙ forget₂ Lat PartOrd = + forget₂ BddLat BddOrd ⋙ forget₂ BddOrd PartOrd := rfl + +lemma forget_SemilatSup_PartOrd_eq_forget_BddOrd_PartOrd : + forget₂ BddLat SemilatSup ⋙ forget₂ SemilatSup PartOrd = + forget₂ BddLat BddOrd ⋙ forget₂ BddOrd PartOrd := rfl + +lemma forget_SemilatInf_PartOrd_eq_forget_BddOrd_PartOrd : + forget₂ BddLat SemilatInf ⋙ forget₂ SemilatInf PartOrd = + forget₂ BddLat BddOrd ⋙ forget₂ BddOrd PartOrd := rfl + +/-- Constructs an equivalence between bounded lattices from an order isomorphism +between them. -/ +@[simps] def iso.mk {α β : BddLat.{u}} (e : α ≃o β) : α ≅ β := +{ hom := e, + inv := e.symm, + hom_inv_id' := by { ext, exact e.symm_apply_apply _ }, + inv_hom_id' := by { ext, exact e.apply_symm_apply _ } } + +/-- `order_dual` as a functor. -/ +@[simps] def dual : BddLat ⥤ BddLat := +{ obj := λ X, of Xᵒᵈ, map := λ X Y, bounded_lattice_hom.dual } + +/-- The equivalence between `BddLat` and itself induced by `order_dual` both ways. -/ +@[simps functor inverse] def dual_equiv : BddLat ≌ BddLat := +equivalence.mk dual dual + (nat_iso.of_components (λ X, iso.mk $ order_iso.dual_dual X) $ λ X Y f, rfl) + (nat_iso.of_components (λ X, iso.mk $ order_iso.dual_dual X) $ λ X Y f, rfl) + +end BddLat + +lemma BddLat_dual_comp_forget_to_BddOrd : + BddLat.dual ⋙ forget₂ BddLat BddOrd = + forget₂ BddLat BddOrd ⋙ BddOrd.dual := rfl + +lemma BddLat_dual_comp_forget_to_Lat : + BddLat.dual ⋙ forget₂ BddLat Lat = + forget₂ BddLat Lat ⋙ Lat.dual := rfl + +lemma BddLat_dual_comp_forget_to_SemilatSup : + BddLat.dual ⋙ forget₂ BddLat SemilatSup = + forget₂ BddLat SemilatInf ⋙ SemilatInf.dual := rfl + +lemma BddLat_dual_comp_forget_to_SemilatInf : + BddLat.dual ⋙ forget₂ BddLat SemilatInf = + forget₂ BddLat SemilatSup ⋙ SemilatSup.dual := rfl diff --git a/src/order/category/BddOrd.lean b/src/order/category/BddOrd.lean new file mode 100644 index 0000000000000..0be06702f6a2f --- /dev/null +++ b/src/order/category/BddOrd.lean @@ -0,0 +1,83 @@ +/- +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 category_theory.category.Bipointed +import order.category.PartOrd +import order.hom.bounded + +/-! +# The category of bounded orders + +This defines `BddOrd`, the category of bounded orders. +-/ + +universes u v + +open category_theory + +/-- The category of bounded orders with monotone functions. -/ +structure BddOrd := +(to_PartOrd : PartOrd) +[is_bounded_order : bounded_order to_PartOrd] + +namespace BddOrd + +instance : has_coe_to_sort BddOrd Type* := induced_category.has_coe_to_sort to_PartOrd +instance (X : BddOrd) : partial_order X := X.to_PartOrd.str + +attribute [instance] BddOrd.is_bounded_order + +/-- Construct a bundled `BddOrd` from a `fintype` `partial_order`. -/ +def of (α : Type*) [partial_order α] [bounded_order α] : BddOrd := ⟨⟨α⟩⟩ + +@[simp] lemma coe_of (α : Type*) [partial_order α] [bounded_order α] : ↥(of α) = α := rfl + +instance : inhabited BddOrd := ⟨of punit⟩ + +instance large_category : large_category.{u} BddOrd := +{ hom := λ X Y, bounded_order_hom X Y, + id := λ X, bounded_order_hom.id X, + comp := λ X Y Z f g, g.comp f, + id_comp' := λ X Y, bounded_order_hom.comp_id, + comp_id' := λ X Y, bounded_order_hom.id_comp, + assoc' := λ W X Y Z _ _ _, bounded_order_hom.comp_assoc _ _ _ } + +instance concrete_category : concrete_category BddOrd := +{ forget := ⟨coe_sort, λ X Y, coe_fn, λ X, rfl, λ X Y Z f g, rfl⟩, + forget_faithful := ⟨λ X Y, by convert fun_like.coe_injective⟩ } + +instance has_forget_to_PartOrd : has_forget₂ BddOrd PartOrd := +{ forget₂ := { obj := λ X, X.to_PartOrd, map := λ X Y, bounded_order_hom.to_order_hom } } + +instance has_forget_to_Bipointed : has_forget₂ BddOrd Bipointed := +{ forget₂ := { obj := λ X, ⟨X, ⊥, ⊤⟩, map := λ X Y f, ⟨f, map_bot f, map_top f⟩ }, + forget_comp := rfl } + +/-- `order_dual` as a functor. -/ +@[simps] def dual : BddOrd ⥤ BddOrd := +{ 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 {α β : BddOrd.{u}} (e : α ≃o β) : α ≅ β := +{ hom := e, + inv := e.symm, + hom_inv_id' := by { ext, exact e.symm_apply_apply _ }, + inv_hom_id' := by { ext, exact e.apply_symm_apply _ } } + +/-- The equivalence between `BddOrd` and itself induced by `order_dual` both ways. -/ +@[simps functor inverse] def dual_equiv : BddOrd ≌ BddOrd := +equivalence.mk dual dual + (nat_iso.of_components (λ X, iso.mk $ order_iso.dual_dual X) $ λ X Y f, rfl) + (nat_iso.of_components (λ X, iso.mk $ order_iso.dual_dual X) $ λ X Y f, rfl) + +end BddOrd + +lemma BddOrd_dual_comp_forget_to_PartOrd : + BddOrd.dual ⋙ forget₂ BddOrd PartOrd = + forget₂ BddOrd PartOrd ⋙ PartOrd.dual := rfl + +lemma BddOrd_dual_comp_forget_to_Bipointed : + BddOrd.dual ⋙ forget₂ BddOrd Bipointed = + forget₂ BddOrd Bipointed ⋙ Bipointed.swap := rfl diff --git a/src/order/category/BoolAlg.lean b/src/order/category/BoolAlg.lean index 6aaaa3a2fab5f..08e38c978beb0 100644 --- a/src/order/category/BoolAlg.lean +++ b/src/order/category/BoolAlg.lean @@ -32,16 +32,16 @@ def of (α : Type*) [boolean_algebra α] : BoolAlg := bundled.of α instance : inhabited BoolAlg := ⟨of punit⟩ -/-- Turn a `BoolAlg` into a `BoundedDistribLattice` by forgetting its complement operation. -/ -def to_BoundedDistribLattice (X : BoolAlg) : BoundedDistribLattice := BoundedDistribLattice.of X +/-- Turn a `BoolAlg` into a `BddDistLat` by forgetting its complement operation. -/ +def to_BddDistLat (X : BoolAlg) : BddDistLat := BddDistLat.of X -@[simp] lemma coe_to_BoundedDistribLattice (X : BoolAlg) : ↥X.to_BoundedDistribLattice = ↥X := rfl +@[simp] lemma coe_to_BddDistLat (X : BoolAlg) : ↥X.to_BddDistLat = ↥X := rfl -instance : large_category.{u} BoolAlg := induced_category.category to_BoundedDistribLattice -instance : concrete_category BoolAlg := induced_category.concrete_category to_BoundedDistribLattice +instance : large_category.{u} BoolAlg := induced_category.category to_BddDistLat +instance : concrete_category BoolAlg := induced_category.concrete_category to_BddDistLat -instance has_forget_to_BoundedDistribLattice : has_forget₂ BoolAlg BoundedDistribLattice := -induced_category.has_forget₂ to_BoundedDistribLattice +instance has_forget_to_BddDistLat : has_forget₂ BoolAlg BddDistLat := +induced_category.has_forget₂ to_BddDistLat section @@ -71,6 +71,6 @@ equivalence.mk dual dual end BoolAlg -lemma BoolAlg_dual_comp_forget_to_BoundedDistribLattice : - BoolAlg.dual ⋙ forget₂ BoolAlg BoundedDistribLattice = - forget₂ BoolAlg BoundedDistribLattice ⋙ BoundedDistribLattice.dual := rfl +lemma BoolAlg_dual_comp_forget_to_BddDistLat : + BoolAlg.dual ⋙ forget₂ BoolAlg BddDistLat = + forget₂ BoolAlg BddDistLat ⋙ BddDistLat.dual := rfl diff --git a/src/order/category/BoundedDistribLattice.lean b/src/order/category/BoundedDistribLattice.lean deleted file mode 100644 index ecac2f1ba4af5..0000000000000 --- a/src/order/category/BoundedDistribLattice.lean +++ /dev/null @@ -1,83 +0,0 @@ -/- -Copyright (c) 2022 Yaël Dillies. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yaël Dillies --/ -import order.category.BoundedLattice -import order.category.DistribLattice - -/-! -# The category of bounded distributive lattices - -This defines `BoundedDistribLattice`, the category of bounded distributive lattices. - -Note that this category is sometimes called [`DistLat`](https://ncatlab.org/nlab/show/DistLat) when -being a lattice is understood to entail having a bottom and a top element. --/ - -universes u - -open category_theory - -/-- The category of bounded distributive lattices with bounded lattice morphisms. -/ -structure BoundedDistribLattice := -(to_DistribLattice : DistribLattice) -[is_bounded_order : bounded_order to_DistribLattice] - -namespace BoundedDistribLattice - -instance : has_coe_to_sort BoundedDistribLattice Type* := ⟨λ X, X.to_DistribLattice⟩ -instance (X : BoundedDistribLattice) : distrib_lattice X := X.to_DistribLattice.str - -attribute [instance] BoundedDistribLattice.is_bounded_order - -/-- Construct a bundled `BoundedDistribLattice` from a `bounded_order` `distrib_lattice`. -/ -def of (α : Type*) [distrib_lattice α] [bounded_order α] : BoundedDistribLattice := ⟨⟨α⟩⟩ - -@[simp] lemma coe_of (α : Type*) [distrib_lattice α] [bounded_order α] : ↥(of α) = α := rfl - -instance : inhabited BoundedDistribLattice := ⟨of punit⟩ - -/-- Turn a `BoundedDistribLattice` into a `BoundedLattice` by forgetting it is distributive. -/ -def to_BoundedLattice (X : BoundedDistribLattice) : BoundedLattice := BoundedLattice.of X - -@[simp] lemma coe_to_BoundedLattice (X : BoundedDistribLattice) : ↥X.to_BoundedLattice = ↥X := rfl - -instance : large_category.{u} BoundedDistribLattice := induced_category.category to_BoundedLattice - -instance : concrete_category BoundedDistribLattice := -induced_category.concrete_category to_BoundedLattice - -instance has_forget_to_DistribLattice : has_forget₂ BoundedDistribLattice DistribLattice := -{ forget₂ := { obj := λ X, ⟨X⟩, map := λ X Y, bounded_lattice_hom.to_lattice_hom } } - -instance has_forget_to_BoundedLattice : has_forget₂ BoundedDistribLattice BoundedLattice := -induced_category.has_forget₂ to_BoundedLattice - -lemma forget_BoundedLattice_Lattice_eq_forget_DistribLattice_Lattice : - forget₂ BoundedDistribLattice BoundedLattice ⋙ forget₂ BoundedLattice Lattice = - forget₂ BoundedDistribLattice DistribLattice ⋙ forget₂ DistribLattice Lattice := rfl - -/-- Constructs an equivalence between bounded distributive lattices from an order isomorphism -between them. -/ -@[simps] def iso.mk {α β : BoundedDistribLattice.{u}} (e : α ≃o β) : α ≅ β := -{ hom := (e : bounded_lattice_hom α β), - inv := (e.symm : bounded_lattice_hom β α), - hom_inv_id' := by { ext, exact e.symm_apply_apply _ }, - inv_hom_id' := by { ext, exact e.apply_symm_apply _ } } - -/-- `order_dual` as a functor. -/ -@[simps] def dual : BoundedDistribLattice ⥤ BoundedDistribLattice := -{ 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 := -equivalence.mk dual dual - (nat_iso.of_components (λ X, iso.mk $ order_iso.dual_dual X) $ λ X Y f, rfl) - (nat_iso.of_components (λ X, iso.mk $ order_iso.dual_dual X) $ λ X Y f, rfl) - -end BoundedDistribLattice - -lemma BoundedDistribLattice_dual_comp_forget_to_DistribLattice : - BoundedDistribLattice.dual ⋙ forget₂ BoundedDistribLattice DistribLattice = - forget₂ BoundedDistribLattice DistribLattice ⋙ DistribLattice.dual := rfl diff --git a/src/order/category/BoundedLattice.lean b/src/order/category/BoundedLattice.lean deleted file mode 100644 index 0201a28d1a191..0000000000000 --- a/src/order/category/BoundedLattice.lean +++ /dev/null @@ -1,125 +0,0 @@ -/- -Copyright (c) 2022 Yaël Dillies. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yaël Dillies --/ -import order.category.BoundedOrder -import order.category.Lattice -import order.category.Semilattice - -/-! -# The category of bounded lattices - -This file defines `BoundedLattice`, the category of bounded lattices. - -In literature, this is sometimes called `Lat`, the category of lattices, because being a lattice is -understood to entail having a bottom and a top element. --/ - -universes u - -open category_theory - -/-- The category of bounded lattices with bounded lattice morphisms. -/ -structure BoundedLattice := -(to_Lattice : Lattice) -[is_bounded_order : bounded_order to_Lattice] - -namespace BoundedLattice - -instance : has_coe_to_sort BoundedLattice Type* := ⟨λ X, X.to_Lattice⟩ -instance (X : BoundedLattice) : lattice X := X.to_Lattice.str - -attribute [instance] BoundedLattice.is_bounded_order - -/-- Construct a bundled `BoundedLattice` from `lattice` + `bounded_order`. -/ -def of (α : Type*) [lattice α] [bounded_order α] : BoundedLattice := ⟨⟨α⟩⟩ - -@[simp] lemma coe_of (α : Type*) [lattice α] [bounded_order α] : ↥(of α) = α := rfl - -instance : inhabited BoundedLattice := ⟨of punit⟩ - -instance : large_category.{u} BoundedLattice := -{ hom := λ X Y, bounded_lattice_hom X Y, - id := λ X, bounded_lattice_hom.id X, - comp := λ X Y Z f g, g.comp f, - id_comp' := λ X Y, bounded_lattice_hom.comp_id, - comp_id' := λ X Y, bounded_lattice_hom.id_comp, - assoc' := λ W X Y Z _ _ _, bounded_lattice_hom.comp_assoc _ _ _ } - -instance : concrete_category BoundedLattice := -{ forget := ⟨coe_sort, λ X Y, coe_fn, λ X, rfl, λ X Y Z f g, rfl⟩, - forget_faithful := ⟨λ X Y, by convert fun_like.coe_injective⟩ } - -instance has_forget_to_BoundedOrder : has_forget₂ BoundedLattice BoundedOrder := -{ forget₂ := { obj := λ X, BoundedOrder.of X, - map := λ X Y, bounded_lattice_hom.to_bounded_order_hom } } - -instance has_forget_to_Lattice : has_forget₂ BoundedLattice Lattice := -{ forget₂ := { obj := λ X, ⟨X⟩, map := λ X Y, bounded_lattice_hom.to_lattice_hom } } - -instance has_forget_to_SemilatticeSup : has_forget₂ BoundedLattice SemilatticeSup := -{ forget₂ := { obj := λ X, ⟨X⟩, map := λ X Y, bounded_lattice_hom.to_sup_bot_hom } } - -instance has_forget_to_SemilatticeInf : has_forget₂ BoundedLattice SemilatticeInf := -{ forget₂ := { obj := λ X, ⟨X⟩, map := λ X Y, bounded_lattice_hom.to_inf_top_hom } } - -@[simp] lemma coe_forget_to_BoundedOrder (X : BoundedLattice) : - ↥((forget₂ BoundedLattice BoundedOrder).obj X) = ↥X := rfl - -@[simp] lemma coe_forget_to_Lattice (X : BoundedLattice) : - ↥((forget₂ BoundedLattice Lattice).obj X) = ↥X := rfl - -@[simp] lemma coe_forget_to_SemilatticeSup (X : BoundedLattice) : - ↥((forget₂ BoundedLattice SemilatticeSup).obj X) = ↥X := rfl - -@[simp] lemma coe_forget_to_SemilatticeInf (X : BoundedLattice) : - ↥((forget₂ BoundedLattice SemilatticeInf).obj X) = ↥X := rfl - -lemma forget_Lattice_PartialOrder_eq_forget_BoundedOrder_PartialOrder : - forget₂ BoundedLattice Lattice ⋙ forget₂ Lattice PartialOrder = - forget₂ BoundedLattice BoundedOrder ⋙ forget₂ BoundedOrder PartialOrder := rfl - -lemma forget_SemilatticeSup_PartialOrder_eq_forget_BoundedOrder_PartialOrder : - forget₂ BoundedLattice SemilatticeSup ⋙ forget₂ SemilatticeSup PartialOrder = - forget₂ BoundedLattice BoundedOrder ⋙ forget₂ BoundedOrder PartialOrder := rfl - -lemma forget_SemilatticeInf_PartialOrder_eq_forget_BoundedOrder_PartialOrder : - forget₂ BoundedLattice SemilatticeInf ⋙ forget₂ SemilatticeInf PartialOrder = - forget₂ BoundedLattice BoundedOrder ⋙ forget₂ BoundedOrder PartialOrder := rfl - -/-- Constructs an equivalence between bounded lattices from an order isomorphism -between them. -/ -@[simps] def iso.mk {α β : BoundedLattice.{u}} (e : α ≃o β) : α ≅ β := -{ hom := e, - inv := e.symm, - hom_inv_id' := by { ext, exact e.symm_apply_apply _ }, - inv_hom_id' := by { ext, exact e.apply_symm_apply _ } } - -/-- `order_dual` as a functor. -/ -@[simps] def dual : BoundedLattice ⥤ BoundedLattice := -{ 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 := -equivalence.mk dual dual - (nat_iso.of_components (λ X, iso.mk $ order_iso.dual_dual X) $ λ X Y f, rfl) - (nat_iso.of_components (λ X, iso.mk $ order_iso.dual_dual X) $ λ X Y f, rfl) - -end BoundedLattice - -lemma BoundedLattice_dual_comp_forget_to_BoundedOrder : - BoundedLattice.dual ⋙ forget₂ BoundedLattice BoundedOrder = - forget₂ BoundedLattice BoundedOrder ⋙ BoundedOrder.dual := rfl - -lemma BoundedLattice_dual_comp_forget_to_Lattice : - BoundedLattice.dual ⋙ forget₂ BoundedLattice Lattice = - forget₂ BoundedLattice Lattice ⋙ Lattice.dual := rfl - -lemma BoundedLattice_dual_comp_forget_to_SemilatticeSup : - BoundedLattice.dual ⋙ forget₂ BoundedLattice SemilatticeSup = - forget₂ BoundedLattice SemilatticeInf ⋙ SemilatticeInf.dual := rfl - -lemma BoundedLattice_dual_comp_forget_to_SemilatticeInf : - BoundedLattice.dual ⋙ forget₂ BoundedLattice SemilatticeInf = - forget₂ BoundedLattice SemilatticeSup ⋙ SemilatticeSup.dual := rfl diff --git a/src/order/category/BoundedOrder.lean b/src/order/category/BoundedOrder.lean deleted file mode 100644 index 8c359e2282b31..0000000000000 --- a/src/order/category/BoundedOrder.lean +++ /dev/null @@ -1,83 +0,0 @@ -/- -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 category_theory.category.Bipointed -import order.category.PartialOrder -import order.hom.bounded - -/-! -# The category of bounded orders - -This defines `BoundedOrder`, the category of bounded orders. --/ - -universes u v - -open category_theory - -/-- The category of bounded orders with monotone functions. -/ -structure BoundedOrder := -(to_PartialOrder : PartialOrder) -[is_bounded_order : bounded_order to_PartialOrder] - -namespace BoundedOrder - -instance : has_coe_to_sort BoundedOrder Type* := induced_category.has_coe_to_sort to_PartialOrder -instance (X : BoundedOrder) : partial_order X := X.to_PartialOrder.str - -attribute [instance] BoundedOrder.is_bounded_order - -/-- Construct a bundled `BoundedOrder` from a `fintype` `partial_order`. -/ -def of (α : Type*) [partial_order α] [bounded_order α] : BoundedOrder := ⟨⟨α⟩⟩ - -@[simp] lemma coe_of (α : Type*) [partial_order α] [bounded_order α] : ↥(of α) = α := rfl - -instance : inhabited BoundedOrder := ⟨of punit⟩ - -instance large_category : large_category.{u} BoundedOrder := -{ hom := λ X Y, bounded_order_hom X Y, - id := λ X, bounded_order_hom.id X, - comp := λ X Y Z f g, g.comp f, - id_comp' := λ X Y, bounded_order_hom.comp_id, - comp_id' := λ X Y, bounded_order_hom.id_comp, - assoc' := λ W X Y Z _ _ _, bounded_order_hom.comp_assoc _ _ _ } - -instance concrete_category : concrete_category BoundedOrder := -{ forget := ⟨coe_sort, λ X Y, coe_fn, λ X, rfl, λ X Y Z f g, rfl⟩, - forget_faithful := ⟨λ X Y, by convert fun_like.coe_injective⟩ } - -instance has_forget_to_PartialOrder : has_forget₂ BoundedOrder PartialOrder := -{ forget₂ := { obj := λ X, X.to_PartialOrder, map := λ X Y, bounded_order_hom.to_order_hom } } - -instance has_forget_to_Bipointed : has_forget₂ BoundedOrder Bipointed := -{ forget₂ := { obj := λ X, ⟨X, ⊥, ⊤⟩, map := λ X Y f, ⟨f, map_bot f, map_top f⟩ }, - forget_comp := rfl } - -/-- `order_dual` as a functor. -/ -@[simps] def dual : BoundedOrder ⥤ BoundedOrder := -{ 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 β) : α ≅ β := -{ hom := e, - inv := e.symm, - hom_inv_id' := by { ext, exact e.symm_apply_apply _ }, - inv_hom_id' := by { ext, exact e.apply_symm_apply _ } } - -/-- The equivalence between `BoundedOrder` and itself induced by `order_dual` both ways. -/ -@[simps functor inverse] def dual_equiv : BoundedOrder ≌ BoundedOrder := -equivalence.mk dual dual - (nat_iso.of_components (λ X, iso.mk $ order_iso.dual_dual X) $ λ X Y f, rfl) - (nat_iso.of_components (λ X, iso.mk $ order_iso.dual_dual X) $ λ X Y f, rfl) - -end BoundedOrder - -lemma BoundedOrder_dual_comp_forget_to_PartialOrder : - BoundedOrder.dual ⋙ forget₂ BoundedOrder PartialOrder = - forget₂ BoundedOrder PartialOrder ⋙ PartialOrder.dual := rfl - -lemma BoundedOrder_dual_comp_forget_to_Bipointed : - BoundedOrder.dual ⋙ forget₂ BoundedOrder Bipointed = - forget₂ BoundedOrder Bipointed ⋙ Bipointed.swap := rfl diff --git a/src/order/category/CompleteLat.lean b/src/order/category/CompleteLat.lean new file mode 100644 index 0000000000000..1649a5cc2e6fc --- /dev/null +++ b/src/order/category/CompleteLat.lean @@ -0,0 +1,68 @@ +/- +Copyright (c) 2022 Yaël Dillies. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies +-/ +import order.category.BddLat +import order.hom.complete_lattice + +/-! +# The category of complete lattices + +This file defines `CompleteLat`, the category of complete lattices. +-/ + +universes u + +open category_theory + +/-- The category of complete lattices. -/ +def CompleteLat := bundled complete_lattice + +namespace CompleteLat + +instance : has_coe_to_sort CompleteLat Type* := bundled.has_coe_to_sort +instance (X : CompleteLat) : complete_lattice X := X.str + +/-- Construct a bundled `CompleteLat` from a `complete_lattice`. -/ +def of (α : Type*) [complete_lattice α] : CompleteLat := bundled.of α + +@[simp] lemma coe_of (α : Type*) [complete_lattice α] : ↥(of α) = α := rfl + +instance : inhabited CompleteLat := ⟨of punit⟩ + +instance : bundled_hom @complete_lattice_hom := +{ to_fun := λ _ _ _ _, coe_fn, + id := @complete_lattice_hom.id, + comp := @complete_lattice_hom.comp, + hom_ext := λ X Y _ _, by exactI fun_like.coe_injective } +instance : large_category.{u} CompleteLat := bundled_hom.category complete_lattice_hom +instance : concrete_category CompleteLat := bundled_hom.concrete_category complete_lattice_hom + +instance has_forget_to_BddLat : has_forget₂ CompleteLat BddLat := +{ forget₂ := { obj := λ X, BddLat.of X, + map := λ X Y, complete_lattice_hom.to_bounded_lattice_hom }, + forget_comp := rfl } + +/-- Constructs an isomorphism of complete lattices from an order isomorphism between them. -/ +@[simps] def iso.mk {α β : CompleteLat.{u}} (e : α ≃o β) : α ≅ β := +{ hom := e, + inv := e.symm, + hom_inv_id' := by { ext, exact e.symm_apply_apply _ }, + inv_hom_id' := by { ext, exact e.apply_symm_apply _ } } + +/-- `order_dual` as a functor. -/ +@[simps] def dual : CompleteLat ⥤ CompleteLat := +{ obj := λ X, of Xᵒᵈ, map := λ X Y, complete_lattice_hom.dual } + +/-- The equivalence between `CompleteLat` and itself induced by `order_dual` both ways. -/ +@[simps functor inverse] def dual_equiv : CompleteLat ≌ CompleteLat := +equivalence.mk dual dual + (nat_iso.of_components (λ X, iso.mk $ order_iso.dual_dual X) $ λ X Y f, rfl) + (nat_iso.of_components (λ X, iso.mk $ order_iso.dual_dual X) $ λ X Y f, rfl) + +end CompleteLat + +lemma CompleteLat_dual_comp_forget_to_BddLat : + CompleteLat.dual ⋙ forget₂ CompleteLat BddLat = + forget₂ CompleteLat BddLat ⋙ BddLat.dual := rfl diff --git a/src/order/category/CompleteLattice.lean b/src/order/category/CompleteLattice.lean deleted file mode 100644 index 03ce1f1839b3a..0000000000000 --- a/src/order/category/CompleteLattice.lean +++ /dev/null @@ -1,68 +0,0 @@ -/- -Copyright (c) 2022 Yaël Dillies. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yaël Dillies --/ -import order.category.BoundedLattice -import order.hom.complete_lattice - -/-! -# The category of complete lattices - -This file defines `CompleteLattice`, the category of complete lattices. --/ - -universes u - -open category_theory - -/-- The category of complete lattices. -/ -def CompleteLattice := bundled complete_lattice - -namespace CompleteLattice - -instance : has_coe_to_sort CompleteLattice Type* := bundled.has_coe_to_sort -instance (X : CompleteLattice) : complete_lattice X := X.str - -/-- Construct a bundled `CompleteLattice` from a `complete_lattice`. -/ -def of (α : Type*) [complete_lattice α] : CompleteLattice := bundled.of α - -@[simp] lemma coe_of (α : Type*) [complete_lattice α] : ↥(of α) = α := rfl - -instance : inhabited CompleteLattice := ⟨of punit⟩ - -instance : bundled_hom @complete_lattice_hom := -{ to_fun := λ _ _ _ _, coe_fn, - id := @complete_lattice_hom.id, - comp := @complete_lattice_hom.comp, - hom_ext := λ X Y _ _, by exactI fun_like.coe_injective } -instance : large_category.{u} CompleteLattice := bundled_hom.category complete_lattice_hom -instance : concrete_category CompleteLattice := bundled_hom.concrete_category complete_lattice_hom - -instance has_forget_to_BoundedLattice : has_forget₂ CompleteLattice BoundedLattice := -{ forget₂ := { obj := λ X, BoundedLattice.of X, - map := λ X Y, complete_lattice_hom.to_bounded_lattice_hom }, - forget_comp := rfl } - -/-- Constructs an isomorphism of complete lattices from an order isomorphism between them. -/ -@[simps] def iso.mk {α β : CompleteLattice.{u}} (e : α ≃o β) : α ≅ β := -{ hom := e, - inv := e.symm, - hom_inv_id' := by { ext, exact e.symm_apply_apply _ }, - inv_hom_id' := by { ext, exact e.apply_symm_apply _ } } - -/-- `order_dual` as a functor. -/ -@[simps] def dual : CompleteLattice ⥤ CompleteLattice := -{ 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 := -equivalence.mk dual dual - (nat_iso.of_components (λ X, iso.mk $ order_iso.dual_dual X) $ λ X Y f, rfl) - (nat_iso.of_components (λ X, iso.mk $ order_iso.dual_dual X) $ λ X Y f, rfl) - -end CompleteLattice - -lemma CompleteLattice_dual_comp_forget_to_BoundedLattice : - CompleteLattice.dual ⋙ forget₂ CompleteLattice BoundedLattice = - forget₂ CompleteLattice BoundedLattice ⋙ BoundedLattice.dual := rfl diff --git a/src/order/category/DistLat.lean b/src/order/category/DistLat.lean new file mode 100644 index 0000000000000..c351ee58c40db --- /dev/null +++ b/src/order/category/DistLat.lean @@ -0,0 +1,65 @@ +/- +Copyright (c) 2022 Yaël Dillies. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies +-/ +import order.category.Lat + +/-! +# The category of distributive lattices + +This file defines `DistLat`, the category of distributive lattices. + +Note that [`DistLat`](https://ncatlab.org/nlab/show/DistLat) in the literature doesn't always +correspond to `DistLat` as we don't require bottom or top elements. Instead, this `DistLat` +corresponds to `BddDistLat`. +-/ + +universes u + +open category_theory + +/-- The category of distributive lattices. -/ +def DistLat := bundled distrib_lattice + +namespace DistLat + +instance : has_coe_to_sort DistLat Type* := bundled.has_coe_to_sort +instance (X : DistLat) : distrib_lattice X := X.str + +/-- Construct a bundled `DistLat` from a `distrib_lattice` underlying type and typeclass. -/ +def of (α : Type*) [distrib_lattice α] : DistLat := bundled.of α + +@[simp] lemma coe_of (α : Type*) [distrib_lattice α] : ↥(of α) = α := rfl + +instance : inhabited DistLat := ⟨of punit⟩ + +instance : bundled_hom.parent_projection @distrib_lattice.to_lattice := ⟨⟩ + +attribute [derive [large_category, concrete_category]] DistLat + +instance has_forget_to_Lat : has_forget₂ DistLat Lat := bundled_hom.forget₂ _ _ + +/-- Constructs an equivalence between distributive lattices from an order isomorphism between them. +-/ +@[simps] def iso.mk {α β : DistLat.{u}} (e : α ≃o β) : α ≅ β := +{ hom := e, + inv := e.symm, + hom_inv_id' := by { ext, exact e.symm_apply_apply _ }, + inv_hom_id' := by { ext, exact e.apply_symm_apply _ } } + +/-- `order_dual` as a functor. -/ +@[simps] def dual : DistLat ⥤ DistLat := +{ obj := λ X, of Xᵒᵈ, map := λ X Y, lattice_hom.dual } + +/-- The equivalence between `DistLat` and itself induced by `order_dual` both ways. -/ +@[simps functor inverse] def dual_equiv : DistLat ≌ DistLat := +equivalence.mk dual dual + (nat_iso.of_components (λ X, iso.mk $ order_iso.dual_dual X) $ λ X Y f, rfl) + (nat_iso.of_components (λ X, iso.mk $ order_iso.dual_dual X) $ λ X Y f, rfl) + +end DistLat + +lemma DistLat_dual_comp_forget_to_Lat : + DistLat.dual ⋙ forget₂ DistLat Lat = + forget₂ DistLat Lat ⋙ Lat.dual := rfl diff --git a/src/order/category/DistribLattice.lean b/src/order/category/DistribLattice.lean deleted file mode 100644 index 7b7220477b2b9..0000000000000 --- a/src/order/category/DistribLattice.lean +++ /dev/null @@ -1,65 +0,0 @@ -/- -Copyright (c) 2022 Yaël Dillies. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yaël Dillies --/ -import order.category.Lattice - -/-! -# The category of distributive lattices - -This file defines `DistribLattice`, the category of distributive lattices. - -Note that [`DistLat`](https://ncatlab.org/nlab/show/DistLat) in the literature doesn't always -correspond to `DistribLattice` as we don't require bottom or top elements. Instead, this `DistLat` -corresponds to `BoundedDistribLattice`. --/ - -universes u - -open category_theory - -/-- The category of distributive lattices. -/ -def DistribLattice := bundled distrib_lattice - -namespace DistribLattice - -instance : has_coe_to_sort DistribLattice Type* := bundled.has_coe_to_sort -instance (X : DistribLattice) : distrib_lattice X := X.str - -/-- Construct a bundled `DistribLattice` from a `distrib_lattice` underlying type and typeclass. -/ -def of (α : Type*) [distrib_lattice α] : DistribLattice := bundled.of α - -@[simp] lemma coe_of (α : Type*) [distrib_lattice α] : ↥(of α) = α := rfl - -instance : inhabited DistribLattice := ⟨of punit⟩ - -instance : bundled_hom.parent_projection @distrib_lattice.to_lattice := ⟨⟩ - -attribute [derive [large_category, concrete_category]] DistribLattice - -instance has_forget_to_Lattice : has_forget₂ DistribLattice Lattice := bundled_hom.forget₂ _ _ - -/-- Constructs an equivalence between distributive lattices from an order isomorphism between them. --/ -@[simps] def iso.mk {α β : DistribLattice.{u}} (e : α ≃o β) : α ≅ β := -{ hom := e, - inv := e.symm, - hom_inv_id' := by { ext, exact e.symm_apply_apply _ }, - inv_hom_id' := by { ext, exact e.apply_symm_apply _ } } - -/-- `order_dual` as a functor. -/ -@[simps] def dual : DistribLattice ⥤ DistribLattice := -{ 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 := -equivalence.mk dual dual - (nat_iso.of_components (λ X, iso.mk $ order_iso.dual_dual X) $ λ X Y f, rfl) - (nat_iso.of_components (λ X, iso.mk $ order_iso.dual_dual X) $ λ X Y f, rfl) - -end DistribLattice - -lemma DistribLattice_dual_comp_forget_to_Lattice : - DistribLattice.dual ⋙ forget₂ DistribLattice Lattice = - forget₂ DistribLattice Lattice ⋙ Lattice.dual := rfl diff --git a/src/order/category/FinBoolAlg.lean b/src/order/category/FinBoolAlg.lean index 0ad67e9e3254c..4da81128cd7cd 100644 --- a/src/order/category/FinBoolAlg.lean +++ b/src/order/category/FinBoolAlg.lean @@ -5,7 +5,7 @@ Authors: Yaël Dillies -/ import data.fintype.powerset import order.category.BoolAlg -import order.category.FinPartialOrder +import order.category.FinPartOrd import order.hom.complete_lattice /-! @@ -60,11 +60,11 @@ instance forget_to_BoolAlg_full : full (forget₂ FinBoolAlg BoolAlg) := induced instance forget_to_BoolAlg_faithful : faithful (forget₂ FinBoolAlg BoolAlg) := induced_category.faithful _ -@[simps] instance has_forget_to_FinPartialOrder : has_forget₂ FinBoolAlg FinPartialOrder := -{ forget₂ := { obj := λ X, FinPartialOrder.of X, map := λ X Y f, +@[simps] instance has_forget_to_FinPartOrd : has_forget₂ FinBoolAlg FinPartOrd := +{ forget₂ := { obj := λ X, FinPartOrd.of X, map := λ X Y f, show order_hom X Y, from ↑(show bounded_lattice_hom X Y, from f) } } -instance forget_to_FinPartialOrder_faithful : faithful (forget₂ FinBoolAlg FinPartialOrder) := +instance forget_to_FinPartOrd_faithful : faithful (forget₂ FinBoolAlg FinPartOrd) := ⟨λ X Y f g h, by { have := congr_arg (coe_fn : _ → X → Y) h, exact fun_like.coe_injective this }⟩ /-- Constructs an equivalence between finite Boolean algebras from an order isomorphism between diff --git a/src/order/category/FinPartOrd.lean b/src/order/category/FinPartOrd.lean new file mode 100644 index 0000000000000..189d468eded93 --- /dev/null +++ b/src/order/category/FinPartOrd.lean @@ -0,0 +1,79 @@ +/- +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 category_theory.Fintype +import order.category.PartOrd + +/-! +# The category of finite partial orders + +This defines `FinPartOrd`, the category of finite partial orders. + +Note: `FinPartOrd` is NOT a subcategory of `BddOrd` because its morphisms do not +preserve `⊥` and `⊤`. + +## TODO + +`FinPartOrd` is equivalent to a small category. +-/ + +universes u v + +open category_theory + +/-- The category of finite partial orders with monotone functions. -/ +structure FinPartOrd := +(to_PartOrd : PartOrd) +[is_fintype : fintype to_PartOrd] + +namespace FinPartOrd + +instance : has_coe_to_sort FinPartOrd Type* := ⟨λ X, X.to_PartOrd⟩ +instance (X : FinPartOrd) : partial_order X := X.to_PartOrd.str +attribute [instance] FinPartOrd.is_fintype + +@[simp] lemma coe_to_PartOrd (X : FinPartOrd) : ↥X.to_PartOrd = ↥X := rfl + +/-- Construct a bundled `FinPartOrd` from `fintype` + `partial_order`. -/ +def of (α : Type*) [partial_order α] [fintype α] : FinPartOrd := ⟨⟨α⟩⟩ + +@[simp] lemma coe_of (α : Type*) [partial_order α] [fintype α] : ↥(of α) = α := rfl + +instance : inhabited FinPartOrd := ⟨of punit⟩ + +instance large_category : large_category FinPartOrd := +induced_category.category FinPartOrd.to_PartOrd + +instance concrete_category : concrete_category FinPartOrd := +induced_category.concrete_category FinPartOrd.to_PartOrd + +instance has_forget_to_PartOrd : has_forget₂ FinPartOrd PartOrd := +induced_category.has_forget₂ FinPartOrd.to_PartOrd + +instance has_forget_to_Fintype : has_forget₂ FinPartOrd Fintype := +{ forget₂ := { obj := λ X, ⟨X⟩, map := λ X Y, coe_fn } } + +/-- Constructs an isomorphism of finite partial orders from an order isomorphism between them. -/ +@[simps] def iso.mk {α β : FinPartOrd.{u}} (e : α ≃o β) : α ≅ β := +{ hom := e, + inv := e.symm, + hom_inv_id' := by { ext, exact e.symm_apply_apply _ }, + inv_hom_id' := by { ext, exact e.apply_symm_apply _ } } + +/-- `order_dual` as a functor. -/ +@[simps] def dual : FinPartOrd ⥤ FinPartOrd := +{ obj := λ X, of Xᵒᵈ, map := λ X Y, order_hom.dual } + +/-- The equivalence between `FinPartOrd` and itself induced by `order_dual` both ways. -/ +@[simps functor inverse] def dual_equiv : FinPartOrd ≌ FinPartOrd := +equivalence.mk dual dual + (nat_iso.of_components (λ X, iso.mk $ order_iso.dual_dual X) $ λ X Y f, rfl) + (nat_iso.of_components (λ X, iso.mk $ order_iso.dual_dual X) $ λ X Y f, rfl) + +end FinPartOrd + +lemma FinPartOrd_dual_comp_forget_to_PartOrd : + FinPartOrd.dual ⋙ forget₂ FinPartOrd PartOrd = + forget₂ FinPartOrd PartOrd ⋙ PartOrd.dual := rfl diff --git a/src/order/category/FinPartialOrder.lean b/src/order/category/FinPartialOrder.lean deleted file mode 100644 index 78b7eae8b2fd9..0000000000000 --- a/src/order/category/FinPartialOrder.lean +++ /dev/null @@ -1,79 +0,0 @@ -/- -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 category_theory.Fintype -import order.category.PartialOrder - -/-! -# The category of finite partial orders - -This defines `FinPartialOrder`, the category of finite partial orders. - -Note: `FinPartialOrder` is NOT a subcategory of `BoundedOrder` because its morphisms do not -preserve `⊥` and `⊤`. - -## TODO - -`FinPartialOrder` is equivalent to a small category. --/ - -universes u v - -open category_theory - -/-- The category of finite partial orders with monotone functions. -/ -structure FinPartialOrder := -(to_PartialOrder : PartialOrder) -[is_fintype : fintype to_PartialOrder] - -namespace FinPartialOrder - -instance : has_coe_to_sort FinPartialOrder Type* := ⟨λ X, X.to_PartialOrder⟩ -instance (X : FinPartialOrder) : partial_order X := X.to_PartialOrder.str -attribute [instance] FinPartialOrder.is_fintype - -@[simp] lemma coe_to_PartialOrder (X : FinPartialOrder) : ↥X.to_PartialOrder = ↥X := rfl - -/-- Construct a bundled `FinPartialOrder` from `fintype` + `partial_order`. -/ -def of (α : Type*) [partial_order α] [fintype α] : FinPartialOrder := ⟨⟨α⟩⟩ - -@[simp] lemma coe_of (α : Type*) [partial_order α] [fintype α] : ↥(of α) = α := rfl - -instance : inhabited FinPartialOrder := ⟨of punit⟩ - -instance large_category : large_category FinPartialOrder := -induced_category.category FinPartialOrder.to_PartialOrder - -instance concrete_category : concrete_category FinPartialOrder := -induced_category.concrete_category FinPartialOrder.to_PartialOrder - -instance has_forget_to_PartialOrder : has_forget₂ FinPartialOrder PartialOrder := -induced_category.has_forget₂ FinPartialOrder.to_PartialOrder - -instance has_forget_to_Fintype : has_forget₂ FinPartialOrder Fintype := -{ forget₂ := { obj := λ X, ⟨X⟩, map := λ X Y, coe_fn } } - -/-- Constructs an isomorphism of finite partial orders from an order isomorphism between them. -/ -@[simps] def iso.mk {α β : FinPartialOrder.{u}} (e : α ≃o β) : α ≅ β := -{ hom := e, - inv := e.symm, - hom_inv_id' := by { ext, exact e.symm_apply_apply _ }, - inv_hom_id' := by { ext, exact e.apply_symm_apply _ } } - -/-- `order_dual` as a functor. -/ -@[simps] def dual : FinPartialOrder ⥤ FinPartialOrder := -{ 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 := -equivalence.mk dual dual - (nat_iso.of_components (λ X, iso.mk $ order_iso.dual_dual X) $ λ X Y f, rfl) - (nat_iso.of_components (λ X, iso.mk $ order_iso.dual_dual X) $ λ X Y f, rfl) - -end FinPartialOrder - -lemma FinPartialOrder_dual_comp_forget_to_PartialOrder : - FinPartialOrder.dual ⋙ forget₂ FinPartialOrder PartialOrder = - forget₂ FinPartialOrder PartialOrder ⋙ PartialOrder.dual := rfl diff --git a/src/order/category/Frame.lean b/src/order/category/Frm.lean similarity index 69% rename from src/order/category/Frame.lean rename to src/order/category/Frm.lean index f50e545168f00..61e67e1efa97e 100644 --- a/src/order/category/Frame.lean +++ b/src/order/category/Frm.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ -import order.category.Lattice +import order.category.Lat import order.hom.complete_lattice import topology.category.CompHaus.basic import topology.sets.opens @@ -11,7 +11,7 @@ import topology.sets.opens /-! # The category of frames -This file defines `Frame`, the category of frames. +This file defines `Frm`, the category of frames. ## References @@ -23,19 +23,19 @@ universes u open category_theory opposite order topological_space /-- The category of frames. -/ -def Frame := bundled frame +def Frm := bundled frame -namespace Frame +namespace Frm -instance : has_coe_to_sort Frame Type* := bundled.has_coe_to_sort -instance (X : Frame) : frame X := X.str +instance : has_coe_to_sort Frm Type* := bundled.has_coe_to_sort +instance (X : Frm) : frame X := X.str -/-- Construct a bundled `Frame` from a `frame`. -/ -def of (α : Type*) [frame α] : Frame := bundled.of α +/-- Construct a bundled `Frm` from a `frame`. -/ +def of (α : Type*) [frame α] : Frm := bundled.of α @[simp] lemma coe_of (α : Type*) [frame α] : ↥(of α) = α := rfl -instance : inhabited Frame := ⟨of punit⟩ +instance : inhabited Frm := ⟨of punit⟩ /-- An abbreviation of `frame_hom` that assumes `frame` instead of the weaker `complete_lattice`. Necessary for the category theory machinery. -/ @@ -47,23 +47,23 @@ instance bundled_hom : bundled_hom hom := λ α β γ [frame α] [frame β] [frame γ], by exactI frame_hom.comp, λ α β [frame α] [frame β], by exactI fun_like.coe_injective⟩ -attribute [derive [large_category, concrete_category]] Frame +attribute [derive [large_category, concrete_category]] Frm -instance has_forget_to_Lattice : has_forget₂ Frame Lattice := +instance has_forget_to_Lat : has_forget₂ Frm Lat := { forget₂ := { obj := λ X, ⟨X⟩, map := λ X Y, frame_hom.to_lattice_hom } } /-- Constructs an isomorphism of frames from an order isomorphism between them. -/ -@[simps] def iso.mk {α β : Frame.{u}} (e : α ≃o β) : α ≅ β := +@[simps] def iso.mk {α β : Frm.{u}} (e : α ≃o β) : α ≅ β := { hom := e, inv := e.symm, hom_inv_id' := by { ext, exact e.symm_apply_apply _ }, inv_hom_id' := by { ext, exact e.apply_symm_apply _ } } -end Frame +end Frm -/-- The forgetful functor from `Topᵒᵖ` to `Frame`. -/ -@[simps] def Top_op_to_Frame : Topᵒᵖ ⥤ Frame := -{ obj := λ X, Frame.of (opens (unop X : Top)), +/-- The forgetful functor from `Topᵒᵖ` to `Frm`. -/ +@[simps] def Top_op_to_Frame : Topᵒᵖ ⥤ Frm := +{ obj := λ X, Frm.of (opens (unop X : Top)), map := λ X Y f, opens.comap $ quiver.hom.unop f, map_id' := λ X, opens.comap_id } diff --git a/src/order/category/HeytAlg.lean b/src/order/category/HeytAlg.lean index 5e3c7af5b84b4..efea6f68b1608 100644 --- a/src/order/category/HeytAlg.lean +++ b/src/order/category/HeytAlg.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ -import order.category.BoundedDistribLattice +import order.category.BddDistLat import order.heyting.hom /-! @@ -41,8 +41,8 @@ instance bundled_hom : bundled_hom heyting_hom := attribute [derive [large_category, concrete_category]] HeytAlg @[simps] -instance has_forget_to_Lattice : has_forget₂ HeytAlg BoundedDistribLattice := -{ forget₂ := { obj := λ X, BoundedDistribLattice.of X, +instance has_forget_to_Lat : has_forget₂ HeytAlg BddDistLat := +{ forget₂ := { obj := λ X, BddDistLat.of X, map := λ X Y f, (f : bounded_lattice_hom X Y) } } /-- Constructs an isomorphism of Heyting algebras from an order isomorphism between them. -/ diff --git a/src/order/category/Lat.lean b/src/order/category/Lat.lean new file mode 100644 index 0000000000000..85a51ca335bc8 --- /dev/null +++ b/src/order/category/Lat.lean @@ -0,0 +1,75 @@ +/- +Copyright (c) 2022 Yaël Dillies. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies +-/ +import order.category.PartOrd +import order.hom.lattice + +/-! +# The category of lattices + +This defines `Lat`, the category of lattices. + +Note that `Lat` doesn't correspond to the literature definition of [`Lat`] +(https://ncatlab.org/nlab/show/Lat) as we don't require bottom or top elements. Instead, `Lat` +corresponds to `BddLat`. + +## TODO + +The free functor from `Lat` to `BddLat` is `X → with_top (with_bot X)`. +-/ + +universes u + +open category_theory + +/-- The category of lattices. -/ +def Lat := bundled lattice + +namespace Lat + +instance : has_coe_to_sort Lat Type* := bundled.has_coe_to_sort +instance (X : Lat) : lattice X := X.str + +/-- Construct a bundled `Lat` from a `lattice`. -/ +def of (α : Type*) [lattice α] : Lat := bundled.of α + +@[simp] lemma coe_of (α : Type*) [lattice α] : ↥(of α) = α := rfl + +instance : inhabited Lat := ⟨of bool⟩ + +instance : bundled_hom @lattice_hom := +{ to_fun := λ _ _ _ _, coe_fn, + id := @lattice_hom.id, + comp := @lattice_hom.comp, + hom_ext := λ X Y _ _, by exactI fun_like.coe_injective } + +instance : large_category.{u} Lat := bundled_hom.category lattice_hom +instance : concrete_category Lat := bundled_hom.concrete_category lattice_hom + +instance has_forget_to_PartOrd : has_forget₂ Lat PartOrd := +{ forget₂ := { obj := λ X, ⟨X⟩, map := λ X Y f, f }, + forget_comp := rfl } + +/-- Constructs an isomorphism of lattices from an order isomorphism between them. -/ +@[simps] def iso.mk {α β : Lat.{u}} (e : α ≃o β) : α ≅ β := +{ hom := e, + inv := e.symm, + hom_inv_id' := by { ext, exact e.symm_apply_apply _ }, + inv_hom_id' := by { ext, exact e.apply_symm_apply _ } } + +/-- `order_dual` as a functor. -/ +@[simps] def dual : Lat ⥤ Lat := { obj := λ X, of Xᵒᵈ, map := λ X Y, lattice_hom.dual } + +/-- The equivalence between `Lat` and itself induced by `order_dual` both ways. -/ +@[simps functor inverse] def dual_equiv : Lat ≌ Lat := +equivalence.mk dual dual + (nat_iso.of_components (λ X, iso.mk $ order_iso.dual_dual X) $ λ X Y f, rfl) + (nat_iso.of_components (λ X, iso.mk $ order_iso.dual_dual X) $ λ X Y f, rfl) + +end Lat + +lemma Lat_dual_comp_forget_to_PartOrd : + Lat.dual ⋙ forget₂ Lat PartOrd = + forget₂ Lat PartOrd ⋙ PartOrd.dual := rfl diff --git a/src/order/category/Lattice.lean b/src/order/category/Lattice.lean deleted file mode 100644 index 3fa08568307f4..0000000000000 --- a/src/order/category/Lattice.lean +++ /dev/null @@ -1,75 +0,0 @@ -/- -Copyright (c) 2022 Yaël Dillies. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yaël Dillies --/ -import order.category.PartialOrder -import order.hom.lattice - -/-! -# The category of lattices - -This defines `Lattice`, the category of lattices. - -Note that `Lattice` doesn't correspond to the literature definition of [`Lat`] -(https://ncatlab.org/nlab/show/Lat) as we don't require bottom or top elements. Instead, `Lat` -corresponds to `BoundedLattice` (not yet in mathlib). - -## TODO - -The free functor from `Lattice` to `BoundedLattice` is `X → with_top (with_bot X)`. --/ - -universes u - -open category_theory - -/-- The category of lattices. -/ -def Lattice := bundled lattice - -namespace Lattice - -instance : has_coe_to_sort Lattice Type* := bundled.has_coe_to_sort -instance (X : Lattice) : lattice X := X.str - -/-- Construct a bundled `Lattice` from a `lattice`. -/ -def of (α : Type*) [lattice α] : Lattice := bundled.of α - -@[simp] lemma coe_of (α : Type*) [lattice α] : ↥(of α) = α := rfl - -instance : inhabited Lattice := ⟨of bool⟩ - -instance : bundled_hom @lattice_hom := -{ to_fun := λ _ _ _ _, coe_fn, - id := @lattice_hom.id, - comp := @lattice_hom.comp, - hom_ext := λ X Y _ _, by exactI fun_like.coe_injective } - -instance : large_category.{u} Lattice := bundled_hom.category lattice_hom -instance : concrete_category Lattice := bundled_hom.concrete_category lattice_hom - -instance has_forget_to_PartialOrder : has_forget₂ Lattice PartialOrder := -{ forget₂ := { obj := λ X, ⟨X⟩, map := λ X Y f, f }, - forget_comp := rfl } - -/-- Constructs an isomorphism of lattices from an order isomorphism between them. -/ -@[simps] def iso.mk {α β : Lattice.{u}} (e : α ≃o β) : α ≅ β := -{ hom := e, - inv := e.symm, - hom_inv_id' := by { ext, exact e.symm_apply_apply _ }, - inv_hom_id' := by { ext, exact e.apply_symm_apply _ } } - -/-- `order_dual` as a functor. -/ -@[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 := -equivalence.mk dual dual - (nat_iso.of_components (λ X, iso.mk $ order_iso.dual_dual X) $ λ X Y f, rfl) - (nat_iso.of_components (λ X, iso.mk $ order_iso.dual_dual X) $ λ X Y f, rfl) - -end Lattice - -lemma Lattice_dual_comp_forget_to_PartialOrder : - Lattice.dual ⋙ forget₂ Lattice PartialOrder = - forget₂ Lattice PartialOrder ⋙ PartialOrder.dual := rfl diff --git a/src/order/category/LinOrd.lean b/src/order/category/LinOrd.lean new file mode 100644 index 0000000000000..f3eefd21eeba9 --- /dev/null +++ b/src/order/category/LinOrd.lean @@ -0,0 +1,64 @@ +/- +Copyright (c) 2020 Johan Commelin. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johan Commelin +-/ + +import order.category.Lat + +/-! +# Category of linear orders + +This defines `LinOrd`, the category of linear orders with monotone maps. +-/ + +open category_theory + +universe u + +/-- The category of linear orders. -/ +def LinOrd := bundled linear_order + +namespace LinOrd + +instance : bundled_hom.parent_projection @linear_order.to_partial_order := ⟨⟩ + +attribute [derive [large_category, concrete_category]] LinOrd + +instance : has_coe_to_sort LinOrd Type* := bundled.has_coe_to_sort + +/-- Construct a bundled `LinOrd` from the underlying type and typeclass. -/ +def of (α : Type*) [linear_order α] : LinOrd := bundled.of α + +@[simp] lemma coe_of (α : Type*) [linear_order α] : ↥(of α) = α := rfl + +instance : inhabited LinOrd := ⟨of punit⟩ + +instance (α : LinOrd) : linear_order α := α.str + +instance has_forget_to_Lat : has_forget₂ LinOrd Lat := +{ forget₂ := { obj := λ X, Lat.of X, + map := λ X Y f, (order_hom_class.to_lattice_hom X Y f : lattice_hom X Y) } } + +/-- Constructs an equivalence between linear orders from an order isomorphism between them. -/ +@[simps] def iso.mk {α β : LinOrd.{u}} (e : α ≃o β) : α ≅ β := +{ hom := e, + inv := e.symm, + hom_inv_id' := by { ext, exact e.symm_apply_apply x }, + inv_hom_id' := by { ext, exact e.apply_symm_apply x } } + +/-- `order_dual` as a functor. -/ +@[simps] def dual : LinOrd ⥤ LinOrd := +{ obj := λ X, of Xᵒᵈ, map := λ X Y, order_hom.dual } + +/-- The equivalence between `LinOrd` and itself induced by `order_dual` both ways. -/ +@[simps functor inverse] def dual_equiv : LinOrd ≌ LinOrd := +equivalence.mk dual dual + (nat_iso.of_components (λ X, iso.mk $ order_iso.dual_dual X) $ λ X Y f, rfl) + (nat_iso.of_components (λ X, iso.mk $ order_iso.dual_dual X) $ λ X Y f, rfl) + +end LinOrd + +lemma LinOrd_dual_comp_forget_to_Lat : + LinOrd.dual ⋙ forget₂ LinOrd Lat = forget₂ LinOrd Lat ⋙ Lat.dual := +rfl diff --git a/src/order/category/LinearOrder.lean b/src/order/category/LinearOrder.lean deleted file mode 100644 index e8d0bda96df25..0000000000000 --- a/src/order/category/LinearOrder.lean +++ /dev/null @@ -1,64 +0,0 @@ -/- -Copyright (c) 2020 Johan Commelin. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Johan Commelin --/ - -import order.category.Lattice - -/-! -# Category of linear orders - -This defines `LinearOrder`, the category of linear orders with monotone maps. --/ - -open category_theory - -universe u - -/-- The category of linear orders. -/ -def LinearOrder := bundled linear_order - -namespace LinearOrder - -instance : bundled_hom.parent_projection @linear_order.to_partial_order := ⟨⟩ - -attribute [derive [large_category, concrete_category]] LinearOrder - -instance : has_coe_to_sort LinearOrder Type* := bundled.has_coe_to_sort - -/-- Construct a bundled `LinearOrder` from the underlying type and typeclass. -/ -def of (α : Type*) [linear_order α] : LinearOrder := bundled.of α - -@[simp] lemma coe_of (α : Type*) [linear_order α] : ↥(of α) = α := rfl - -instance : inhabited LinearOrder := ⟨of punit⟩ - -instance (α : LinearOrder) : linear_order α := α.str - -instance has_forget_to_Lattice : has_forget₂ LinearOrder Lattice := -{ forget₂ := { obj := λ X, Lattice.of X, - map := λ X Y f, (order_hom_class.to_lattice_hom X Y f : lattice_hom X Y) } } - -/-- Constructs an equivalence between linear orders from an order isomorphism between them. -/ -@[simps] def iso.mk {α β : LinearOrder.{u}} (e : α ≃o β) : α ≅ β := -{ hom := e, - inv := e.symm, - hom_inv_id' := by { ext, exact e.symm_apply_apply x }, - inv_hom_id' := by { ext, exact e.apply_symm_apply x } } - -/-- `order_dual` as a functor. -/ -@[simps] def dual : LinearOrder ⥤ LinearOrder := -{ 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 := -equivalence.mk dual dual - (nat_iso.of_components (λ X, iso.mk $ order_iso.dual_dual X) $ λ X Y f, rfl) - (nat_iso.of_components (λ X, iso.mk $ order_iso.dual_dual X) $ λ X Y f, rfl) - -end LinearOrder - -lemma LinearOrder_dual_comp_forget_to_Lattice : - LinearOrder.dual ⋙ forget₂ LinearOrder Lattice = forget₂ LinearOrder Lattice ⋙ Lattice.dual := -rfl diff --git a/src/order/category/NonemptyFinLinOrd.lean b/src/order/category/NonemptyFinLinOrd.lean index 95a42944534e9..af7ab122b152c 100644 --- a/src/order/category/NonemptyFinLinOrd.lean +++ b/src/order/category/NonemptyFinLinOrd.lean @@ -5,7 +5,7 @@ Authors: Johan Commelin -/ import data.fintype.order import data.set.finite -import order.category.LinearOrder +import order.category.LinOrd import category_theory.limits.shapes.images import category_theory.limits.shapes.regular_mono @@ -62,7 +62,7 @@ instance : inhabited NonemptyFinLinOrd := ⟨of punit⟩ instance (α : NonemptyFinLinOrd) : nonempty_fin_lin_ord α := α.str -instance has_forget_to_LinearOrder : has_forget₂ NonemptyFinLinOrd LinearOrder := +instance has_forget_to_LinOrd : has_forget₂ NonemptyFinLinOrd LinOrd := bundled_hom.forget₂ _ _ /-- Constructs an equivalence between nonempty finite linear orders from an order isomorphism @@ -77,7 +77,7 @@ between them. -/ @[simps] def dual : NonemptyFinLinOrd ⥤ NonemptyFinLinOrd := { obj := λ X, of Xᵒᵈ, map := λ X Y, order_hom.dual } -/-- The equivalence between `FinPartialOrder` and itself induced by `order_dual` both ways. -/ +/-- The equivalence between `FinPartOrd` and itself induced by `order_dual` both ways. -/ @[simps functor inverse] def dual_equiv : NonemptyFinLinOrd ≌ NonemptyFinLinOrd := equivalence.mk dual dual (nat_iso.of_components (λ X, iso.mk $ order_iso.dual_dual X) $ λ X Y f, rfl) @@ -176,6 +176,6 @@ end⟩ end NonemptyFinLinOrd -lemma NonemptyFinLinOrd_dual_comp_forget_to_LinearOrder : - NonemptyFinLinOrd.dual ⋙ forget₂ NonemptyFinLinOrd LinearOrder = - forget₂ NonemptyFinLinOrd LinearOrder ⋙ LinearOrder.dual := rfl +lemma NonemptyFinLinOrd_dual_comp_forget_to_LinOrd : + NonemptyFinLinOrd.dual ⋙ forget₂ NonemptyFinLinOrd LinOrd = + forget₂ NonemptyFinLinOrd LinOrd ⋙ LinOrd.dual := rfl diff --git a/src/order/category/PartialOrder.lean b/src/order/category/PartOrd.lean similarity index 52% rename from src/order/category/PartialOrder.lean rename to src/order/category/PartOrd.lean index 793f355aa681f..b4baeee3989f7 100644 --- a/src/order/category/PartialOrder.lean +++ b/src/order/category/PartOrd.lean @@ -4,12 +4,12 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin -/ import order.antisymmetrization -import order.category.Preorder +import order.category.Preord /-! # Category of partial orders -This defines `PartialOrder`, the category of partial orders with monotone maps. +This defines `PartOrd`, the category of partial orders with monotone maps. -/ open category_theory @@ -17,63 +17,63 @@ open category_theory universe u /-- The category of partially ordered types. -/ -def PartialOrder := bundled partial_order +def PartOrd := bundled partial_order -namespace PartialOrder +namespace PartOrd instance : bundled_hom.parent_projection @partial_order.to_preorder := ⟨⟩ -attribute [derive [large_category, concrete_category]] PartialOrder +attribute [derive [large_category, concrete_category]] PartOrd -instance : has_coe_to_sort PartialOrder Type* := bundled.has_coe_to_sort +instance : has_coe_to_sort PartOrd Type* := bundled.has_coe_to_sort -/-- Construct a bundled PartialOrder from the underlying type and typeclass. -/ -def of (α : Type*) [partial_order α] : PartialOrder := bundled.of α +/-- Construct a bundled PartOrd from the underlying type and typeclass. -/ +def of (α : Type*) [partial_order α] : PartOrd := bundled.of α @[simp] lemma coe_of (α : Type*) [partial_order α] : ↥(of α) = α := rfl -instance : inhabited PartialOrder := ⟨of punit⟩ +instance : inhabited PartOrd := ⟨of punit⟩ -instance (α : PartialOrder) : partial_order α := α.str +instance (α : PartOrd) : partial_order α := α.str -instance has_forget_to_Preorder : has_forget₂ PartialOrder Preorder := bundled_hom.forget₂ _ _ +instance has_forget_to_Preord : has_forget₂ PartOrd Preord := bundled_hom.forget₂ _ _ /-- Constructs an equivalence between partial orders from an order isomorphism between them. -/ -@[simps] def iso.mk {α β : PartialOrder.{u}} (e : α ≃o β) : α ≅ β := +@[simps] def iso.mk {α β : PartOrd.{u}} (e : α ≃o β) : α ≅ β := { hom := e, inv := e.symm, hom_inv_id' := by { ext, exact e.symm_apply_apply x }, inv_hom_id' := by { ext, exact e.apply_symm_apply x } } /-- `order_dual` as a functor. -/ -@[simps] def dual : PartialOrder ⥤ PartialOrder := +@[simps] def dual : PartOrd ⥤ PartOrd := { 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 := +/-- The equivalence between `PartOrd` and itself induced by `order_dual` both ways. -/ +@[simps functor inverse] def dual_equiv : PartOrd ≌ PartOrd := equivalence.mk dual dual (nat_iso.of_components (λ X, iso.mk $ order_iso.dual_dual X) $ λ X Y f, rfl) (nat_iso.of_components (λ X, iso.mk $ order_iso.dual_dual X) $ λ X Y f, rfl) -end PartialOrder +end PartOrd -lemma PartialOrder_dual_comp_forget_to_Preorder : - PartialOrder.dual ⋙ forget₂ PartialOrder Preorder = - forget₂ PartialOrder Preorder ⋙ Preorder.dual := rfl +lemma PartOrd_dual_comp_forget_to_Preord : + PartOrd.dual ⋙ forget₂ PartOrd Preord = + forget₂ PartOrd Preord ⋙ Preord.dual := rfl /-- `antisymmetrization` as a functor. It is the free functor. -/ -def Preorder_to_PartialOrder : Preorder.{u} ⥤ PartialOrder := -{ obj := λ X, PartialOrder.of (antisymmetrization X (≤)), +def Preord_to_PartOrd : Preord.{u} ⥤ PartOrd := +{ obj := λ X, PartOrd.of (antisymmetrization X (≤)), map := λ X Y f, f.antisymmetrization, map_id' := λ X, by { ext, exact quotient.induction_on' x (λ x, quotient.map'_mk' _ (λ a b, id) _) }, map_comp' := λ X Y Z f g, by { ext, exact quotient.induction_on' x (λ x, order_hom.antisymmetrization_apply_mk _ _) } } -/-- `Preorder_to_PartialOrder` is left adjoint to the forgetful functor, meaning it is the free -functor from `Preorder` to `PartialOrder`. -/ -def Preorder_to_PartialOrder_forget_adjunction : - Preorder_to_PartialOrder.{u} ⊣ forget₂ PartialOrder Preorder := +/-- `Preord_to_PartOrd` is left adjoint to the forgetful functor, meaning it is the free +functor from `Preord` to `PartOrd`. -/ +def Preord_to_PartOrd_forget_adjunction : + Preord_to_PartOrd.{u} ⊣ forget₂ PartOrd Preord := adjunction.mk_of_hom_equiv { hom_equiv := λ X Y, { to_fun := λ f, ⟨f ∘ to_antisymmetrization (≤), f.mono.comp to_antisymmetrization_mono⟩, @@ -85,9 +85,9 @@ adjunction.mk_of_hom_equiv order_hom.ext _ _ $ funext $ λ x, quotient.induction_on' x $ λ x, rfl, hom_equiv_naturality_right' := λ X Y Z f g, order_hom.ext _ _ $ funext $ λ x, rfl } -/-- `Preorder_to_PartialOrder` and `order_dual` commute. -/ -@[simps] def Preorder_to_PartialOrder_comp_to_dual_iso_to_dual_comp_Preorder_to_PartialOrder : - (Preorder_to_PartialOrder.{u} ⋙ PartialOrder.dual) ≅ - (Preorder.dual ⋙ Preorder_to_PartialOrder) := -nat_iso.of_components (λ X, PartialOrder.iso.mk $ order_iso.dual_antisymmetrization _) $ +/-- `Preord_to_PartOrd` and `order_dual` commute. -/ +@[simps] def Preord_to_PartOrd_comp_to_dual_iso_to_dual_comp_Preord_to_PartOrd : + (Preord_to_PartOrd.{u} ⋙ PartOrd.dual) ≅ + (Preord.dual ⋙ Preord_to_PartOrd) := +nat_iso.of_components (λ X, PartOrd.iso.mk $ order_iso.dual_antisymmetrization _) $ λ X Y f, order_hom.ext _ _ $ funext $ λ x, quotient.induction_on' x $ λ x, rfl diff --git a/src/order/category/Preorder.lean b/src/order/category/Preord.lean similarity index 62% rename from src/order/category/Preorder.lean rename to src/order/category/Preord.lean index 75cb06958c741..ee40020817599 100644 --- a/src/order/category/Preorder.lean +++ b/src/order/category/Preord.lean @@ -11,7 +11,7 @@ import order.hom.basic /-! # Category of preorders -This defines `Preorder`, the category of preorders with monotone maps. +This defines `Preord`, the category of preorders with monotone maps. -/ universe u @@ -19,9 +19,9 @@ universe u open category_theory /-- The category of preorders. -/ -def Preorder := bundled preorder +def Preord := bundled preorder -namespace Preorder +namespace Preord instance : bundled_hom @order_hom := { to_fun := @order_hom.to_fun, @@ -29,51 +29,51 @@ instance : bundled_hom @order_hom := comp := @order_hom.comp, hom_ext := @order_hom.ext } -attribute [derive [large_category, concrete_category]] Preorder +attribute [derive [large_category, concrete_category]] Preord -instance : has_coe_to_sort Preorder Type* := bundled.has_coe_to_sort +instance : has_coe_to_sort Preord Type* := bundled.has_coe_to_sort -/-- Construct a bundled Preorder from the underlying type and typeclass. -/ -def of (α : Type*) [preorder α] : Preorder := bundled.of α +/-- Construct a bundled Preord from the underlying type and typeclass. -/ +def of (α : Type*) [preorder α] : Preord := bundled.of α @[simp] lemma coe_of (α : Type*) [preorder α] : ↥(of α) = α := rfl -instance : inhabited Preorder := ⟨of punit⟩ +instance : inhabited Preord := ⟨of punit⟩ -instance (α : Preorder) : preorder α := α.str +instance (α : Preord) : preorder α := α.str /-- Constructs an equivalence between preorders from an order isomorphism between them. -/ -@[simps] def iso.mk {α β : Preorder.{u}} (e : α ≃o β) : α ≅ β := +@[simps] def iso.mk {α β : Preord.{u}} (e : α ≃o β) : α ≅ β := { hom := e, inv := e.symm, hom_inv_id' := by { ext, exact e.symm_apply_apply x }, inv_hom_id' := by { ext, exact e.apply_symm_apply x } } /-- `order_dual` as a functor. -/ -@[simps] def dual : Preorder ⥤ Preorder := +@[simps] def dual : Preord ⥤ Preord := { 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 := +/-- The equivalence between `Preord` and itself induced by `order_dual` both ways. -/ +@[simps functor inverse] def dual_equiv : Preord ≌ Preord := equivalence.mk dual dual (nat_iso.of_components (λ X, iso.mk $ order_iso.dual_dual X) $ λ X Y f, rfl) (nat_iso.of_components (λ X, iso.mk $ order_iso.dual_dual X) $ λ X Y f, rfl) -end Preorder +end Preord /-- -The embedding of `Preorder` into `Cat`. +The embedding of `Preord` into `Cat`. -/ @[simps] -def Preorder_to_Cat : Preorder.{u} ⥤ Cat := +def Preord_to_Cat : Preord.{u} ⥤ Cat := { obj := λ X, Cat.of X.1, map := λ X Y f, f.monotone.functor, map_id' := λ X, begin apply category_theory.functor.ext, tidy end, map_comp' := λ X Y Z f g, begin apply category_theory.functor.ext, tidy end } -instance : faithful Preorder_to_Cat.{u} := +instance : faithful Preord_to_Cat.{u} := { map_injective' := λ X Y f g h, begin ext x, exact functor.congr_obj h x end } -instance : full Preorder_to_Cat.{u} := +instance : full Preord_to_Cat.{u} := { preimage := λ X Y f, ⟨f.obj, f.monotone⟩, witness' := λ X Y f, begin apply category_theory.functor.ext, tidy end } diff --git a/src/order/category/Semilat.lean b/src/order/category/Semilat.lean new file mode 100644 index 0000000000000..44b058fcc2e7f --- /dev/null +++ b/src/order/category/Semilat.lean @@ -0,0 +1,148 @@ +/- +Copyright (c) 2022 Yaël Dillies. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies +-/ +import order.category.PartOrd +import order.hom.lattice + +/-! +# The categories of semilattices + +This defines `SemilatSup` and `SemilatInf`, the categories of sup-semilattices with a bottom +element and inf-semilattices with a top element. + +## References + +* [nLab, *semilattice*](https://ncatlab.org/nlab/show/semilattice) +-/ + +universes u +open category_theory + +/-- The category of sup-semilattices with a bottom element. -/ +structure SemilatSup : Type.{u+1} := +(X : Type.{u}) +[is_semilattice_sup : semilattice_sup X] +[is_order_bot : order_bot X] + +/-- The category of inf-semilattices with a top element. -/ +structure SemilatInf : Type.{u+1} := +(X : Type.{u}) +[is_semilattice_inf : semilattice_inf X] +[is_order_top : order_top X] + +attribute [protected] SemilatSup.X SemilatInf.X + +namespace SemilatSup + +instance : has_coe_to_sort SemilatSup Type* := ⟨SemilatSup.X⟩ +attribute [instance] is_semilattice_sup is_order_bot + +/-- Construct a bundled `SemilatSup` from a `semilattice_sup`. -/ +def of (α : Type*) [semilattice_sup α] [order_bot α] : SemilatSup := ⟨α⟩ + +@[simp] lemma coe_of (α : Type*) [semilattice_sup α] [order_bot α] : ↥(of α) = α := rfl + +instance : inhabited SemilatSup := ⟨of punit⟩ + +instance : large_category.{u} SemilatSup := +{ hom := λ X Y, sup_bot_hom X Y, + id := λ X, sup_bot_hom.id X, + comp := λ X Y Z f g, g.comp f, + id_comp' := λ X Y, sup_bot_hom.comp_id, + comp_id' := λ X Y, sup_bot_hom.id_comp, + assoc' := λ W X Y Z _ _ _, sup_bot_hom.comp_assoc _ _ _ } + +instance : concrete_category SemilatSup := +{ forget := { obj := SemilatSup.X, map := λ X Y, coe_fn }, + forget_faithful := ⟨λ X Y, fun_like.coe_injective⟩ } + +instance has_forget_to_PartOrd : has_forget₂ SemilatSup PartOrd := +{ forget₂ := { obj := λ X, ⟨X⟩, map := λ X Y f, f } } + +@[simp] lemma coe_forget_to_PartOrd (X : SemilatSup) : + ↥((forget₂ SemilatSup PartOrd).obj X) = ↥X := rfl + +end SemilatSup + +namespace SemilatInf + +instance : has_coe_to_sort SemilatInf Type* := ⟨SemilatInf.X⟩ + +attribute [instance] is_semilattice_inf is_order_top + +/-- Construct a bundled `SemilatInf` from a `semilattice_inf`. -/ +def of (α : Type*) [semilattice_inf α] [order_top α] : SemilatInf := ⟨α⟩ + +@[simp] lemma coe_of (α : Type*) [semilattice_inf α] [order_top α] : ↥(of α) = α := rfl + +instance : inhabited SemilatInf := ⟨of punit⟩ + +instance : large_category.{u} SemilatInf := +{ hom := λ X Y, inf_top_hom X Y, + id := λ X, inf_top_hom.id X, + comp := λ X Y Z f g, g.comp f, + id_comp' := λ X Y, inf_top_hom.comp_id, + comp_id' := λ X Y, inf_top_hom.id_comp, + assoc' := λ W X Y Z _ _ _, inf_top_hom.comp_assoc _ _ _ } + +instance : concrete_category SemilatInf := +{ forget := { obj := SemilatInf.X, map := λ X Y, coe_fn }, + forget_faithful := ⟨λ X Y, fun_like.coe_injective⟩ } + +instance has_forget_to_PartOrd : has_forget₂ SemilatInf PartOrd := +{ forget₂ := { obj := λ X, ⟨X⟩, map := λ X Y f, f } } + +@[simp] lemma coe_forget_to_PartOrd (X : SemilatInf) : + ↥((forget₂ SemilatInf PartOrd).obj X) = ↥X := rfl + +end SemilatInf + +/-! ### Order dual -/ + +namespace SemilatSup + +/-- Constructs an isomorphism of lattices from an order isomorphism between them. -/ +@[simps] def iso.mk {α β : SemilatSup.{u}} (e : α ≃o β) : α ≅ β := +{ hom := e, + inv := e.symm, + hom_inv_id' := by { ext, exact e.symm_apply_apply _ }, + inv_hom_id' := by { ext, exact e.apply_symm_apply _ } } + +/-- `order_dual` as a functor. -/ +@[simps] def dual : SemilatSup ⥤ SemilatInf := +{ obj := λ X, SemilatInf.of Xᵒᵈ, map := λ X Y, sup_bot_hom.dual } + +end SemilatSup + +namespace SemilatInf + +/-- Constructs an isomorphism of lattices from an order isomorphism between them. -/ +@[simps] def iso.mk {α β : SemilatInf.{u}} (e : α ≃o β) : α ≅ β := +{ hom := e, + inv := e.symm, + hom_inv_id' := by { ext, exact e.symm_apply_apply _ }, + inv_hom_id' := by { ext, exact e.apply_symm_apply _ } } + +/-- `order_dual` as a functor. -/ +@[simps] def dual : SemilatInf ⥤ SemilatSup := +{ obj := λ X, SemilatSup.of Xᵒᵈ, map := λ X Y, inf_top_hom.dual } + +end SemilatInf + +/-- The equivalence between `SemilatSup` and `SemilatInf` induced by `order_dual` both ways. +-/ +@[simps functor inverse] +def SemilatSup_equiv_SemilatInf : SemilatSup ≌ SemilatInf := +equivalence.mk SemilatSup.dual SemilatInf.dual + (nat_iso.of_components (λ X, SemilatSup.iso.mk $ order_iso.dual_dual X) $ λ X Y f, rfl) + (nat_iso.of_components (λ X, SemilatInf.iso.mk $ order_iso.dual_dual X) $ λ X Y f, rfl) + +lemma SemilatSup_dual_comp_forget_to_PartOrd : + SemilatSup.dual ⋙ forget₂ SemilatInf PartOrd = + forget₂ SemilatSup PartOrd ⋙ PartOrd.dual := rfl + +lemma SemilatInf_dual_comp_forget_to_PartOrd : + SemilatInf.dual ⋙ forget₂ SemilatSup PartOrd = + forget₂ SemilatInf PartOrd ⋙ PartOrd.dual := rfl diff --git a/src/order/category/Semilattice.lean b/src/order/category/Semilattice.lean deleted file mode 100644 index 234a24b254b47..0000000000000 --- a/src/order/category/Semilattice.lean +++ /dev/null @@ -1,148 +0,0 @@ -/- -Copyright (c) 2022 Yaël Dillies. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yaël Dillies --/ -import order.category.PartialOrder -import order.hom.lattice - -/-! -# The categories of semilattices - -This defines `SemilatticeSup` and `SemilatticeInf`, the categories of sup-semilattices with a bottom -element and inf-semilattices with a top element. - -## References - -* [nLab, *semilattice*](https://ncatlab.org/nlab/show/semilattice) --/ - -universes u -open category_theory - -/-- The category of sup-semilattices with a bottom element. -/ -structure SemilatticeSup : Type.{u+1} := -(X : Type.{u}) -[is_semilattice_sup : semilattice_sup X] -[is_order_bot : order_bot X] - -/-- The category of inf-semilattices with a top element. -/ -structure SemilatticeInf : Type.{u+1} := -(X : Type.{u}) -[is_semilattice_inf : semilattice_inf X] -[is_order_top : order_top X] - -attribute [protected] SemilatticeSup.X SemilatticeInf.X - -namespace SemilatticeSup - -instance : has_coe_to_sort SemilatticeSup Type* := ⟨SemilatticeSup.X⟩ -attribute [instance] is_semilattice_sup is_order_bot - -/-- Construct a bundled `SemilatticeSup` from a `semilattice_sup`. -/ -def of (α : Type*) [semilattice_sup α] [order_bot α] : SemilatticeSup := ⟨α⟩ - -@[simp] lemma coe_of (α : Type*) [semilattice_sup α] [order_bot α] : ↥(of α) = α := rfl - -instance : inhabited SemilatticeSup := ⟨of punit⟩ - -instance : large_category.{u} SemilatticeSup := -{ hom := λ X Y, sup_bot_hom X Y, - id := λ X, sup_bot_hom.id X, - comp := λ X Y Z f g, g.comp f, - id_comp' := λ X Y, sup_bot_hom.comp_id, - comp_id' := λ X Y, sup_bot_hom.id_comp, - assoc' := λ W X Y Z _ _ _, sup_bot_hom.comp_assoc _ _ _ } - -instance : concrete_category SemilatticeSup := -{ forget := { obj := SemilatticeSup.X, map := λ X Y, coe_fn }, - forget_faithful := ⟨λ X Y, fun_like.coe_injective⟩ } - -instance has_forget_to_PartialOrder : has_forget₂ SemilatticeSup PartialOrder := -{ forget₂ := { obj := λ X, ⟨X⟩, map := λ X Y f, f } } - -@[simp] lemma coe_forget_to_PartialOrder (X : SemilatticeSup) : - ↥((forget₂ SemilatticeSup PartialOrder).obj X) = ↥X := rfl - -end SemilatticeSup - -namespace SemilatticeInf - -instance : has_coe_to_sort SemilatticeInf Type* := ⟨SemilatticeInf.X⟩ - -attribute [instance] is_semilattice_inf is_order_top - -/-- Construct a bundled `SemilatticeInf` from a `semilattice_inf`. -/ -def of (α : Type*) [semilattice_inf α] [order_top α] : SemilatticeInf := ⟨α⟩ - -@[simp] lemma coe_of (α : Type*) [semilattice_inf α] [order_top α] : ↥(of α) = α := rfl - -instance : inhabited SemilatticeInf := ⟨of punit⟩ - -instance : large_category.{u} SemilatticeInf := -{ hom := λ X Y, inf_top_hom X Y, - id := λ X, inf_top_hom.id X, - comp := λ X Y Z f g, g.comp f, - id_comp' := λ X Y, inf_top_hom.comp_id, - comp_id' := λ X Y, inf_top_hom.id_comp, - assoc' := λ W X Y Z _ _ _, inf_top_hom.comp_assoc _ _ _ } - -instance : concrete_category SemilatticeInf := -{ forget := { obj := SemilatticeInf.X, map := λ X Y, coe_fn }, - forget_faithful := ⟨λ X Y, fun_like.coe_injective⟩ } - -instance has_forget_to_PartialOrder : has_forget₂ SemilatticeInf PartialOrder := -{ forget₂ := { obj := λ X, ⟨X⟩, map := λ X Y f, f } } - -@[simp] lemma coe_forget_to_PartialOrder (X : SemilatticeInf) : - ↥((forget₂ SemilatticeInf PartialOrder).obj X) = ↥X := rfl - -end SemilatticeInf - -/-! ### Order dual -/ - -namespace SemilatticeSup - -/-- Constructs an isomorphism of lattices from an order isomorphism between them. -/ -@[simps] def iso.mk {α β : SemilatticeSup.{u}} (e : α ≃o β) : α ≅ β := -{ hom := e, - inv := e.symm, - hom_inv_id' := by { ext, exact e.symm_apply_apply _ }, - inv_hom_id' := by { ext, exact e.apply_symm_apply _ } } - -/-- `order_dual` as a functor. -/ -@[simps] def dual : SemilatticeSup ⥤ SemilatticeInf := -{ obj := λ X, SemilatticeInf.of Xᵒᵈ, map := λ X Y, sup_bot_hom.dual } - -end SemilatticeSup - -namespace SemilatticeInf - -/-- Constructs an isomorphism of lattices from an order isomorphism between them. -/ -@[simps] def iso.mk {α β : SemilatticeInf.{u}} (e : α ≃o β) : α ≅ β := -{ hom := e, - inv := e.symm, - hom_inv_id' := by { ext, exact e.symm_apply_apply _ }, - inv_hom_id' := by { ext, exact e.apply_symm_apply _ } } - -/-- `order_dual` as a functor. -/ -@[simps] def dual : SemilatticeInf ⥤ SemilatticeSup := -{ obj := λ X, SemilatticeSup.of Xᵒᵈ, map := λ X Y, inf_top_hom.dual } - -end SemilatticeInf - -/-- The equivalence between `SemilatticeSup` and `SemilatticeInf` induced by `order_dual` both ways. --/ -@[simps functor inverse] -def SemilatticeSup_equiv_SemilatticeInf : SemilatticeSup ≌ SemilatticeInf := -equivalence.mk SemilatticeSup.dual SemilatticeInf.dual - (nat_iso.of_components (λ X, SemilatticeSup.iso.mk $ order_iso.dual_dual X) $ λ X Y f, rfl) - (nat_iso.of_components (λ X, SemilatticeInf.iso.mk $ order_iso.dual_dual X) $ λ X Y f, rfl) - -lemma SemilatticeSup_dual_comp_forget_to_PartialOrder : - SemilatticeSup.dual ⋙ forget₂ SemilatticeInf PartialOrder = - forget₂ SemilatticeSup PartialOrder ⋙ PartialOrder.dual := rfl - -lemma SemilatticeInf_dual_comp_forget_to_PartialOrder : - SemilatticeInf.dual ⋙ forget₂ SemilatticeSup PartialOrder = - forget₂ SemilatticeInf PartialOrder ⋙ PartialOrder.dual := rfl diff --git a/src/order/complete_boolean_algebra.lean b/src/order/complete_boolean_algebra.lean index dd590f754bc04..4108c483a0eeb 100644 --- a/src/order/complete_boolean_algebra.lean +++ b/src/order/complete_boolean_algebra.lean @@ -18,7 +18,7 @@ distributive Boolean algebras. ## Typeclasses -* `order.frame`: Frame: A complete lattice whose `⊓` distributes over `⨆`. +* `order.frame`: Frm: A complete lattice whose `⊓` distributes over `⨆`. * `order.coframe`: Coframe: A complete lattice whose `⊔` distributes over `⨅`. * `complete_distrib_lattice`: Completely distributive lattices: A complete lattice whose `⊓` and `⊔` distribute over `⨆` and `⨅` respectively. diff --git a/src/order/hom/complete_lattice.lean b/src/order/hom/complete_lattice.lean index 354e1d03d6106..548321502b4f4 100644 --- a/src/order/hom/complete_lattice.lean +++ b/src/order/hom/complete_lattice.lean @@ -21,7 +21,7 @@ be satisfied by itself and all stricter types. * `Sup_hom`: Maps which preserve `⨆`. * `Inf_hom`: Maps which preserve `⨅`. -* `frame_hom`: Frame homomorphisms. Maps which preserve `⨆`, `⊓` and `⊤`. +* `frame_hom`: Frm homomorphisms. Maps which preserve `⨆`, `⊓` and `⊤`. * `complete_lattice_hom`: Complete lattice homomorphisms. Maps which preserve `⨆` and `⨅`. ## Typeclasses @@ -37,7 +37,7 @@ be satisfied by itself and all stricter types. ## TODO -Frame homs are Heyting homs. +Frm homs are Heyting homs. -/ open function order_dual set @@ -365,7 +365,7 @@ instance : order_top (Inf_hom α β) := ⟨⊤, λ f a, le_top⟩ end Inf_hom -/-! ### Frame homomorphisms -/ +/-! ### Frm homomorphisms -/ namespace frame_hom variables [complete_lattice α] [complete_lattice β] [complete_lattice γ] [complete_lattice δ] diff --git a/src/topology/category/Locale.lean b/src/topology/category/Locale.lean index 85a3223a11f35..795f926e4ebc8 100644 --- a/src/topology/category/Locale.lean +++ b/src/topology/category/Locale.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ -import order.category.Frame +import order.category.Frm /-! # The category of locales @@ -16,7 +16,7 @@ universes u open category_theory opposite order topological_space /-- The category of locales. -/ -@[derive large_category] def Locale := Frameᵒᵖ +@[derive large_category] def Locale := Frmᵒᵖ namespace Locale @@ -24,7 +24,7 @@ instance : has_coe_to_sort Locale Type* := ⟨λ X, X.unop⟩ instance (X : Locale) : frame X := X.unop.str /-- Construct a bundled `Locale` from a `frame`. -/ -def of (α : Type*) [frame α] : Locale := op $ Frame.of α +def of (α : Type*) [frame α] : Locale := op $ Frm.of α @[simp] lemma coe_of (α : Type*) [frame α] : ↥(of α) = α := rfl From 6876fa15e3158ff3e4a4e2af1fb6e1945c6e8803 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Tue, 28 Mar 2023 23:45:18 +0000 Subject: [PATCH 0713/1291] refactor(category_theory/shift): improve automation (#18670) --- src/category_theory/differential_object.lean | 61 +- src/category_theory/graded_object.lean | 13 +- src/category_theory/shift.lean | 436 ----------- src/category_theory/shift/basic.lean | 733 ++++++++++++++++++ src/category_theory/triangulated/basic.lean | 31 +- .../triangulated/pretriangulated.lean | 84 +- src/category_theory/triangulated/rotate.lean | 270 +------ 7 files changed, 843 insertions(+), 785 deletions(-) delete mode 100644 src/category_theory/shift.lean create mode 100644 src/category_theory/shift/basic.lean diff --git a/src/category_theory/differential_object.lean b/src/category_theory/differential_object.lean index 96de5d3b8a33d..345d26ac9dd25 100644 --- a/src/category_theory/differential_object.lean +++ b/src/category_theory/differential_object.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ import data.int.basic -import category_theory.shift +import category_theory.shift.basic import category_theory.concrete_category.basic /-! @@ -220,52 +220,67 @@ def shift_functor (n : ℤ) : differential_object C ⥤ differential_object C := ←functor.map_comp_assoc, X.d_squared, functor.map_zero, zero_comp] }, map := λ X Y f, { f := f.f⟦n⟧', - comm' := by { dsimp, rw [category.assoc, shift_comm_hom_comp, ← functor.map_comp_assoc, - f.comm, functor.map_comp_assoc], }, }, + comm' := begin + dsimp, + erw [category.assoc, shift_comm_hom_comp, ← functor.map_comp_assoc, f.comm, + functor.map_comp_assoc], + refl, + end, }, map_id' := by { intros X, ext1, dsimp, rw functor.map_id }, map_comp' := by { intros X Y Z f g, ext1, dsimp, rw functor.map_comp } } -local attribute [simp] eq_to_hom_map -local attribute [reducible] discrete.add_monoidal shift_comm - /-- The shift functor on `differential_object C` is additive. -/ @[simps] def shift_functor_add (m n : ℤ) : shift_functor C (m + n) ≅ shift_functor C m ⋙ shift_functor C n := begin refine nat_iso.of_components (λ X, mk_iso (shift_add X.X _ _) _) _, { dsimp, - -- This is just `simp, simp [eq_to_hom_map]`. - simp_rw [category.assoc, obj_μ_inv_app, μ_inv_hom_app_assoc, functor.map_comp, obj_μ_app, - category.assoc, μ_naturality_assoc, μ_inv_hom_app_assoc, obj_μ_inv_app, category.assoc, - μ_naturalityₗ_assoc, μ_inv_hom_app_assoc, μ_inv_naturalityᵣ_assoc], - simp only [eq_to_hom_map, eq_to_hom_app, eq_to_iso.hom, eq_to_hom_trans_assoc, - eq_to_iso.inv], }, + rw [← cancel_epi ((shift_functor_add C m n).inv.app X.X)], + simp only [category.assoc, iso.inv_hom_id_app_assoc], + erw [← nat_trans.naturality_assoc], + dsimp, + simp only [functor.map_comp, category.assoc, + shift_functor_comm_hom_app_comp_shift_shift_functor_add_hom_app 1 m n X.X, + iso.inv_hom_id_app_assoc], }, { intros X Y f, ext, dsimp, exact nat_trans.naturality _ _ } end -local attribute [reducible] endofunctor_monoidal_category - section -local attribute [instance] endofunctor_monoidal_category /-- The shift by zero is naturally isomorphic to the identity. -/ @[simps] -def shift_ε : 𝟭 (differential_object C) ≅ shift_functor C 0 := +def shift_zero : shift_functor C 0 ≅ 𝟭 (differential_object C) := begin - refine nat_iso.of_components (λ X, mk_iso ((shift_monoidal_functor C ℤ).ε_iso.app X.X) _) _, - { dsimp, simp, dsimp, simp }, - { introv, ext, dsimp, simp } + refine nat_iso.of_components (λ X, mk_iso ((shift_functor_zero C ℤ).app X.X) _) _, + { erw [← nat_trans.naturality], + dsimp, + simp only [shift_functor_zero_hom_app_shift, category.assoc], }, + { tidy, }, end end -local attribute [simp] eq_to_hom_map - instance : has_shift (differential_object C) ℤ := has_shift_mk _ _ { F := shift_functor C, - ε := shift_ε C, - μ := λ m n, (shift_functor_add C m n).symm } + zero := shift_zero C, + add := shift_functor_add C, + assoc_hom_app := λ m₁ m₂ m₃ X, begin + ext1, + convert shift_functor_add_assoc_hom_app m₁ m₂ m₃ X.X, + dsimp [shift_functor_add'], + simpa, + end, + zero_add_hom_app := λ n X, begin + ext1, + convert shift_functor_add_zero_add_hom_app n X.X, + simpa, + end, + add_zero_hom_app := λ n X, begin + ext1, + convert shift_functor_add_add_zero_hom_app n X.X, + simpa, + end, } end differential_object diff --git a/src/category_theory/graded_object.lean b/src/category_theory/graded_object.lean index 80fe7b9126dc9..2ad314d225879 100644 --- a/src/category_theory/graded_object.lean +++ b/src/category_theory/graded_object.lean @@ -5,7 +5,7 @@ Authors: Scott Morrison -/ import algebra.group_power.lemmas import category_theory.pi.basic -import category_theory.shift +import category_theory.shift.basic import category_theory.concrete_category.basic /-! @@ -106,11 +106,12 @@ instance has_shift {β : Type*} [add_comm_group β] (s : β) : has_shift (graded_object_with_shift s C) ℤ := has_shift_mk _ _ { F := λ n, comap (λ _, C) $ λ (b : β), b + n • s, - ε := (comap_id β (λ _, C)).symm ≪≫ (comap_eq C (by { ext, simp })), - μ := λ m n, comap_comp _ _ _ ≪≫ comap_eq C (by { ext, simp [add_zsmul, add_comm] }), - left_unitality := by { introv, ext, dsimp, simpa }, - right_unitality := by { introv, ext, dsimp, simpa }, - associativity := by { introv, ext, dsimp, simp } } + zero := comap_eq C (by { ext, simp }) ≪≫ comap_id β (λ _, C), + add := λ m n, comap_eq C (by { ext, simp [add_zsmul, add_comm], }) ≪≫ + (comap_comp _ _ _).symm, + assoc_hom_app := λ m₁ m₂ m₃ X, by { ext, dsimp, simp, }, + zero_add_hom_app := λ n X, by { ext, dsimp, simpa, }, + add_zero_hom_app := λ n X, by { ext, dsimp, simpa, }, } @[simp] lemma shift_functor_obj_apply {β : Type*} [add_comm_group β] (s : β) (X : β → C) (t : β) (n : ℤ) : diff --git a/src/category_theory/shift.lean b/src/category_theory/shift.lean deleted file mode 100644 index 51d671b5b7f8c..0000000000000 --- a/src/category_theory/shift.lean +++ /dev/null @@ -1,436 +0,0 @@ -/- -Copyright (c) 2020 Scott Morrison. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Scott Morrison, Johan Commelin, Andrew Yang --/ -import category_theory.limits.preserves.shapes.zero -import category_theory.monoidal.End -import category_theory.monoidal.discrete - -/-! -# Shift - -A `shift` on a category `C` indexed by a monoid `A` is nothing more than a monoidal functor -from `A` to `C ⥤ C`. A typical example to keep in mind might be the category of -complexes `⋯ → C_{n-1} → C_n → C_{n+1} → ⋯`. It has a shift indexed by `ℤ`, where we assign to -each `n : ℤ` the functor `C ⥤ C` that re-indexes the terms, so the degree `i` term of `shift n C` -would be the degree `i+n`-th term of `C`. - -## Main definitions -* `has_shift`: A typeclass asserting the existence of a shift functor. -* `shift_equiv`: When the indexing monoid is a group, then the functor indexed by `n` and `-n` forms - an self-equivalence of `C`. -* `shift_comm`: When the indexing monoid is commutative, then shifts commute as well. - -## Implementation Notes - -Many of the definitions in this file are marked as an `abbreviation` so that the simp lemmas in -`category_theory/monoidal/End` can apply. - --/ -namespace category_theory - -noncomputable theory - -universes v u - -variables (C : Type u) (A : Type*) [category.{v} C] - -local attribute [instance] endofunctor_monoidal_category - -section eq_to_hom - -variables {A C} - -variables [add_monoid A] (F : monoidal_functor (discrete A) (C ⥤ C)) - - @[simp, reassoc] lemma eq_to_hom_μ_app {i j i' j' : A} (h₁ : i = i') (h₂ : j = j') (X : C) : - eq_to_hom (by rw [h₁, h₂] : (F.obj ⟨i⟩ ⊗ F.obj ⟨j⟩).obj X = - (F.obj ⟨i'⟩ ⊗ F.obj ⟨j'⟩).obj X) ≫ (F.μ ⟨i'⟩ ⟨j'⟩).app X = - (F.μ ⟨i⟩ ⟨j⟩).app X ≫ eq_to_hom (by rw [h₁, h₂]) := - by { cases h₁, cases h₂, rw [eq_to_hom_refl, eq_to_hom_refl, category.id_comp, category.comp_id] } - - @[simp, reassoc] lemma μ_inv_app_eq_to_hom {i j i' j' : A} (h₁ : i = i') (h₂ : j = j') (X : C) : - inv ((F.μ ⟨i⟩ ⟨j⟩).app X) ≫ eq_to_hom (by rw [h₁, h₂]) = - eq_to_hom (by rw [h₁, h₂]) ≫ inv ((F.μ ⟨i'⟩ ⟨j'⟩).app X) := - by { cases h₁, cases h₂, rw [eq_to_hom_refl, eq_to_hom_refl, category.id_comp, category.comp_id] } - -end eq_to_hom - -variables {A C} - -/-- A monoidal functor from a group `A` into `C ⥤ C` induces -a self-equivalence of `C` for each `n : A`. -/ -@[simps functor inverse unit_iso_hom unit_iso_inv counit_iso_hom counit_iso_inv] -def add_neg_equiv [add_group A] (F : monoidal_functor (discrete A) (C ⥤ C)) (n : A) : C ≌ C := -equiv_of_tensor_iso_unit F ⟨n⟩ ⟨(-n : A)⟩ - (discrete.eq_to_iso (add_neg_self n)) (discrete.eq_to_iso (neg_add_self n)) - (subsingleton.elim _ _) - -section defs - -variables (A C) [add_monoid A] - -/-- A category has a shift indexed by an additive monoid `A` -if there is a monoidal functor from `A` to `C ⥤ C`. -/ -class has_shift (C : Type u) (A : Type*) [category.{v} C] [add_monoid A] := -(shift : monoidal_functor (discrete A) (C ⥤ C)) - -/-- A helper structure to construct the shift functor `(discrete A) ⥤ (C ⥤ C)`. -/ -@[nolint has_nonempty_instance] -structure shift_mk_core := -(F : A → (C ⥤ C)) -(ε : 𝟭 C ≅ F 0) -(μ : Π n m : A, F n ⋙ F m ≅ F (n + m)) -(associativity : ∀ (m₁ m₂ m₃ : A) (X : C), - (F m₃).map ((μ m₁ m₂).hom.app X) ≫ (μ (m₁ + m₂) m₃).hom.app X ≫ - eq_to_hom (by { congr' 2, exact add_assoc _ _ _ }) = - (μ m₂ m₃).hom.app ((F m₁).obj X) ≫ (μ m₁ (m₂ + m₃)).hom.app X . obviously) -(left_unitality : ∀ (n : A) (X : C), - (F n).map (ε.hom.app X) ≫ (μ 0 n).hom.app X = - eq_to_hom (by { dsimp, rw zero_add }) . obviously) -(right_unitality : ∀ (n : A) (X : C), - ε.hom.app ((F n).obj X) ≫ (μ n 0).hom.app X = - eq_to_hom (by { dsimp, rw add_zero }) . obviously) - -section -local attribute [simp] eq_to_hom_map -local attribute [reducible] endofunctor_monoidal_category discrete.add_monoidal - -/-- Constructs a `has_shift C A` instance from `shift_mk_core`. -/ -@[simps] -def has_shift_mk (h : shift_mk_core C A) : has_shift C A := -⟨{ ε := h.ε.hom, - μ := λ m n, (h.μ m.as n.as).hom, - μ_natural' := by { rintros ⟨X⟩ ⟨Y⟩ ⟨X'⟩ ⟨Y'⟩ ⟨⟨⟨rfl⟩⟩⟩ ⟨⟨⟨rfl⟩⟩⟩, ext, - dsimp, simp, dsimp, simp }, - associativity' := by { introv, ext, dsimp, simpa using h.associativity _ _ _ _, }, - left_unitality' := - by { rintro ⟨X⟩, ext, dsimp, rw [category.id_comp, ← category.assoc, h.left_unitality], simp }, - right_unitality' := - by { rintro ⟨X⟩, ext, dsimp, rw [functor.map_id, category.comp_id, - ← category.assoc, h.right_unitality], simp }, - ..(discrete.functor h.F) }⟩ - -end - -variables [has_shift C A] - -/-- The monoidal functor from `A` to `C ⥤ C` given a `has_shift` instance. -/ -def shift_monoidal_functor : monoidal_functor (discrete A) (C ⥤ C) := has_shift.shift - -variable {A} - -/-- The shift autoequivalence, moving objects and morphisms 'up'. -/ -abbreviation shift_functor (i : A) : C ⥤ C := (shift_monoidal_functor C A).obj ⟨i⟩ - -/-- Shifting by `i + j` is the same as shifting by `i` and then shifting by `j`. -/ -abbreviation shift_functor_add (i j : A) : - shift_functor C (i + j) ≅ shift_functor C i ⋙ shift_functor C j := -((shift_monoidal_functor C A).μ_iso ⟨i⟩ ⟨j⟩).symm - -variables (A) - -/-- Shifting by zero is the identity functor. -/ -abbreviation shift_functor_zero : shift_functor C (0 : A) ≅ 𝟭 C := -(shift_monoidal_functor C A).ε_iso.symm - --- Any better notational suggestions? -notation X`⟦`n`⟧`:20 := (shift_functor _ n).obj X -notation f`⟦`n`⟧'`:80 := (shift_functor _ n).map f - -end defs - -section add_monoid - -variables {C A} [add_monoid A] [has_shift C A] (X Y : C) (f : X ⟶ Y) - -@[simp] lemma has_shift.shift_obj_obj (n : A) (X : C) : (has_shift.shift.obj ⟨n⟩).obj X = X⟦n⟧ := -rfl - -/-- Shifting by `i + j` is the same as shifting by `i` and then shifting by `j`. -/ -abbreviation shift_add (i j : A) : X⟦i + j⟧ ≅ X⟦i⟧⟦j⟧ := (shift_functor_add C i j).app _ - -@[reassoc] lemma shift_add_hom_comp_eq_to_hom₁ (i i' j : A) (h : i = i') : - (shift_add X i j).hom ≫ eq_to_hom (by rw h) = eq_to_hom (by rw h) ≫ (shift_add X i' j).hom := -by { cases h, rw [eq_to_hom_refl, eq_to_hom_refl, category.id_comp, category.comp_id] } - -@[reassoc] lemma shift_add_hom_comp_eq_to_hom₂ (i j j' : A) (h : j = j') : - (shift_add X i j).hom ≫ eq_to_hom (by rw h) = eq_to_hom (by rw h) ≫ (shift_add X i j').hom := -by { cases h, rw [eq_to_hom_refl, eq_to_hom_refl, category.id_comp, category.comp_id] } - -@[reassoc] lemma shift_add_hom_comp_eq_to_hom₁₂ (i j i' j' : A) (h₁ : i = i') (h₂ : j = j') : - (shift_add X i j).hom ≫ eq_to_hom (by rw [h₁, h₂]) = - eq_to_hom (by rw [h₁, h₂]) ≫ (shift_add X i' j').hom := -by { cases h₁, cases h₂, rw [eq_to_hom_refl, eq_to_hom_refl, category.id_comp, category.comp_id] } - -@[reassoc] lemma eq_to_hom_comp_shift_add_inv₁ (i i' j : A) (h : i = i') : - eq_to_hom (by rw h) ≫ (shift_add X i' j).inv = (shift_add X i j).inv ≫ eq_to_hom (by rw h) := -by rw [iso.comp_inv_eq, category.assoc, iso.eq_inv_comp, shift_add_hom_comp_eq_to_hom₁] - -@[reassoc] lemma eq_to_hom_comp_shift_add_inv₂ (i j j' : A) (h : j = j') : - eq_to_hom (by rw h) ≫ (shift_add X i j').inv = (shift_add X i j).inv ≫ eq_to_hom (by rw h) := -by rw [iso.comp_inv_eq, category.assoc, iso.eq_inv_comp, shift_add_hom_comp_eq_to_hom₂] - -@[reassoc] lemma eq_to_hom_comp_shift_add_inv₁₂ (i j i' j' : A) (h₁ : i = i') (h₂ : j = j') : - eq_to_hom (by rw [h₁, h₂]) ≫ (shift_add X i' j').inv = - (shift_add X i j).inv ≫ eq_to_hom (by rw [h₁, h₂]) := -by rw [iso.comp_inv_eq, category.assoc, iso.eq_inv_comp, shift_add_hom_comp_eq_to_hom₁₂] - -lemma shift_shift' (i j : A) : - f⟦i⟧'⟦j⟧' = (shift_add X i j).inv ≫ f⟦i + j⟧' ≫ (shift_add Y i j).hom := -by { symmetry, apply nat_iso.naturality_1 } - -variables (A) - -/-- Shifting by zero is the identity functor. -/ -abbreviation shift_zero : - X⟦0⟧ ≅ X := (shift_functor_zero C A).app _ - -lemma shift_zero' : - f⟦(0 : A)⟧' = (shift_zero A X).hom ≫ f ≫ (shift_zero A Y).inv := -by { symmetry, apply nat_iso.naturality_2 } - -end add_monoid - -section add_group - -variables (C) {A} [add_group A] [has_shift C A] -variables (X Y : C) (f : X ⟶ Y) - -/-- Shifting by `i` is an equivalence. -/ -instance (i : A) : is_equivalence (shift_functor C i) := -begin - change is_equivalence (add_neg_equiv (shift_monoidal_functor C A) i).functor, - apply_instance, -end - -@[simp] lemma shift_functor_inv (i : A) : - (shift_functor C i).inv = shift_functor C (-i) := -rfl - -/-- Shifting by `i` and then shifting by `-i` is the identity. -/ -abbreviation shift_functor_comp_shift_functor_neg (i : A) : - shift_functor C i ⋙ shift_functor C (-i) ≅ 𝟭 C := -unit_of_tensor_iso_unit (shift_monoidal_functor C A) ⟨i⟩ ⟨(-i : A)⟩ - (discrete.eq_to_iso (add_neg_self i)) - -/-- Shifting by `-i` and then shifting by `i` is the identity. -/ -abbreviation shift_functor_neg_comp_shift_functor (i : A) : - shift_functor C (-i) ⋙ shift_functor C i ≅ 𝟭 C := -unit_of_tensor_iso_unit (shift_monoidal_functor C A) ⟨(-i : A)⟩ ⟨i⟩ - (discrete.eq_to_iso (neg_add_self i)) - -section - -variables (C) - -/-- Shifting by `n` is a faithful functor. -/ -instance shift_functor_faithful (i : A) : faithful (shift_functor C i) := -faithful.of_comp_iso (shift_functor_comp_shift_functor_neg C i) - -/-- Shifting by `n` is a full functor. -/ -instance shift_functor_full (i : A) : full (shift_functor C i) := -begin - haveI : full (shift_functor C i ⋙ shift_functor C (-i)) := - full.of_iso (shift_functor_comp_shift_functor_neg C i).symm, - exact full.of_comp_faithful _ (shift_functor C (-i)) -end - -/-- Shifting by `n` is an essentially surjective functor. -/ -instance shift_functor_ess_surj (i : A) : ess_surj (shift_functor C i) := -{ mem_ess_image := λ Y, ⟨Y⟦-i⟧, ⟨(shift_functor_neg_comp_shift_functor C i).app Y⟩⟩ } - -end - -variables {C} - -/-- Shifting by `i` and then shifting by `-i` is the identity. -/ -abbreviation shift_shift_neg (i : A) : X⟦i⟧⟦-i⟧ ≅ X := -(shift_functor_comp_shift_functor_neg C i).app _ - -/-- Shifting by `-i` and then shifting by `i` is the identity. -/ -abbreviation shift_neg_shift (i : A) : X⟦-i⟧⟦i⟧ ≅ X := -(shift_functor_neg_comp_shift_functor C i).app _ - -variables {X Y} - -lemma shift_shift_neg' (i : A) : - f⟦i⟧'⟦-i⟧' = (shift_shift_neg X i).hom ≫ f ≫ (shift_shift_neg Y i).inv := -by { symmetry, apply nat_iso.naturality_2 } - -lemma shift_neg_shift' (i : A) : - f⟦-i⟧'⟦i⟧' = (shift_neg_shift X i).hom ≫ f ≫ (shift_neg_shift Y i).inv := -by { symmetry, apply nat_iso.naturality_2 } - -lemma shift_equiv_triangle (n : A) (X : C) : - (shift_shift_neg X n).inv⟦n⟧' ≫ (shift_neg_shift (X⟦n⟧) n).hom = 𝟙 (X⟦n⟧) := -(add_neg_equiv (shift_monoidal_functor C A) n).functor_unit_iso_comp X - -section -local attribute [reducible] discrete.add_monoidal - -lemma shift_shift_neg_hom_shift (n : A) (X : C) : - (shift_shift_neg X n).hom ⟦n⟧' = (shift_neg_shift (X⟦n⟧) n).hom := -begin - -- This is just `simp, simp [eq_to_hom_map]`. - simp only [iso.app_hom, unit_of_tensor_iso_unit_hom_app, eq_to_iso.hom, functor.map_comp, - obj_μ_app, eq_to_iso.inv, obj_ε_inv_app, μ_naturalityₗ_assoc, category.assoc, - μ_inv_hom_app_assoc, ε_inv_app_obj, μ_naturalityᵣ_assoc], - simp only [eq_to_hom_map, eq_to_hom_app, eq_to_hom_trans], -end - -end - -lemma shift_shift_neg_inv_shift (n : A) (X : C) : - (shift_shift_neg X n).inv ⟦n⟧' = (shift_neg_shift (X⟦n⟧) n).inv := -by { ext, rw [← shift_shift_neg_hom_shift, ← functor.map_comp, iso.hom_inv_id, functor.map_id] } - -@[simp] -lemma shift_shift_neg_shift_eq (n : A) (X : C) : - (shift_functor C n).map_iso (shift_shift_neg X n) = shift_neg_shift (X⟦n⟧) n := -category_theory.iso.ext $ shift_shift_neg_hom_shift _ _ - -variables (C) - -/-- Shifting by `n` and shifting by `-n` forms an equivalence. -/ -@[simps] -def shift_equiv (n : A) : C ≌ C := -{ functor := shift_functor C n, - inverse := shift_functor C (-n), - ..(add_neg_equiv (shift_monoidal_functor C A) n) } - -variable {C} - -open category_theory.limits - -variables [has_zero_morphisms C] - -lemma shift_zero_eq_zero (X Y : C) (n : A) : (0 : X ⟶ Y)⟦n⟧' = (0 : X⟦n⟧ ⟶ Y⟦n⟧) := -category_theory.functor.map_zero _ _ _ - -end add_group - -section add_comm_monoid - -variables {C A} [add_comm_monoid A] [has_shift C A] -variables (X Y : C) (f : X ⟶ Y) - -/-- When shifts are indexed by an additive commutative monoid, then shifts commute. -/ -def shift_comm (i j : A) : X⟦i⟧⟦j⟧ ≅ X⟦j⟧⟦i⟧ := -(shift_add X i j).symm ≪≫ ((shift_monoidal_functor C A).to_functor.map_iso - (discrete.eq_to_iso $ add_comm i j : (⟨i+j⟩ : discrete A) ≅ ⟨j+i⟩)).app X ≪≫ shift_add X j i - -@[simp] lemma shift_comm_symm (i j : A) : (shift_comm X i j).symm = shift_comm X j i := -begin - ext, dsimp [shift_comm], simpa [eq_to_hom_map], -end - -variables {X Y} - -/-- When shifts are indexed by an additive commutative monoid, then shifts commute. -/ -lemma shift_comm' (i j : A) : - f⟦i⟧'⟦j⟧' = (shift_comm _ _ _).hom ≫ f⟦j⟧'⟦i⟧' ≫ (shift_comm _ _ _).hom := -begin - -- This is just `simp, simp [eq_to_hom_map]`. - simp only [shift_comm, iso.trans_hom, iso.symm_hom, iso.app_inv, iso.symm_inv, - monoidal_functor.μ_iso_hom, iso.app_hom, functor.map_iso_hom, eq_to_iso.hom, μ_naturality_assoc, - nat_trans.naturality_assoc, nat_trans.naturality, functor.comp_map, category.assoc, - μ_inv_hom_app_assoc], - simp only [eq_to_hom_map, eq_to_hom_app, eq_to_hom_trans_assoc, eq_to_hom_refl, category.id_comp, - μ_hom_inv_app_assoc], -end - -@[reassoc] lemma shift_comm_hom_comp (i j : A) : - (shift_comm X i j).hom ≫ f⟦j⟧'⟦i⟧' = f⟦i⟧'⟦j⟧' ≫ (shift_comm Y i j).hom := -by rw [shift_comm', ← shift_comm_symm, iso.symm_hom, iso.inv_hom_id_assoc] - -end add_comm_monoid - -variables {D : Type*} [category D] [add_monoid A] [has_shift D A] -variables (F : C ⥤ D) [full F] [faithful F] - -section -local attribute [reducible] discrete.add_monoidal - -/-- Given a family of endomorphisms of `C` which are interwined by a fully faithful `F : C ⥤ D` -with shift functors on `D`, we can promote that family to shift functors on `C`. -/ -def has_shift_of_fully_faithful - (s : A → C ⥤ C) (i : ∀ i, s i ⋙ F ≅ F ⋙ shift_functor D i) : has_shift C A := -has_shift_mk C A -{ F := s, - ε := nat_iso_of_comp_fully_faithful F - (calc 𝟭 C ⋙ F ≅ F : functor.left_unitor _ - ... ≅ F ⋙ 𝟭 D : (functor.right_unitor _).symm - ... ≅ F ⋙ shift_functor D (0 : A) : - iso_whisker_left F (shift_functor_zero D A).symm - ... ≅ s 0 ⋙ F : (i 0).symm), - μ := λ a b, nat_iso_of_comp_fully_faithful F - (calc (s a ⋙ s b) ⋙ F ≅ s a ⋙ s b ⋙ F : functor.associator _ _ _ - ... ≅ s a ⋙ F ⋙ shift_functor D b : iso_whisker_left _ (i b) - ... ≅ (s a ⋙ F) ⋙ shift_functor D b : (functor.associator _ _ _).symm - ... ≅ (F ⋙ shift_functor D a) ⋙ shift_functor D b : iso_whisker_right (i a) _ - ... ≅ F ⋙ shift_functor D a ⋙ shift_functor D b : functor.associator _ _ _ - ... ≅ F ⋙ shift_functor D (a + b) : - iso_whisker_left _ (shift_functor_add D a b).symm - ... ≅ s (a + b) ⋙ F : (i (a + b)).symm), - associativity := begin - intros, apply F.map_injective, dsimp, - simp only [category.comp_id, category.id_comp, category.assoc, - category_theory.functor.map_comp, functor.image_preimage, - eq_to_hom_map, iso.inv_hom_id_app_assoc], - erw (i m₃).hom.naturality_assoc, - congr' 1, - dsimp, - simp only [eq_to_iso.inv, eq_to_hom_app, eq_to_hom_map, obj_μ_app, μ_naturality_assoc, - category.assoc, category_theory.functor.map_comp, functor.image_preimage], - congr' 3, - dsimp, - simp only [←(shift_functor D m₃).map_comp_assoc, iso.inv_hom_id_app], - erw [(shift_functor D m₃).map_id, category.id_comp], - erw [((shift_monoidal_functor D A).μ_iso ⟨m₁ + m₂⟩ ⟨m₃⟩).inv_hom_id_app_assoc], - congr' 1, - have := dcongr_arg (λ a, (i a).inv.app X) (add_assoc m₁ m₂ m₃), - dsimp at this, - simp [this], - end, - left_unitality := begin - intros, apply F.map_injective, dsimp, - simp only [category.comp_id, category.id_comp, category.assoc, category_theory.functor.map_comp, - eq_to_hom_app, eq_to_hom_map, functor.image_preimage], - erw (i n).hom.naturality_assoc, - dsimp, - simp only [eq_to_iso.inv, eq_to_hom_app, category.assoc, category_theory.functor.map_comp, - eq_to_hom_map, obj_ε_app, functor.image_preimage], - simp only [←(shift_functor D n).map_comp_assoc, iso.inv_hom_id_app], - dsimp, - simp only [category.id_comp, μ_inv_hom_app_assoc, category_theory.functor.map_id], - have := dcongr_arg (λ a, (i a).inv.app X) (zero_add n), - dsimp at this, - simp [this], - end, - right_unitality := begin - intros, apply F.map_injective, dsimp, - simp only [category.comp_id, category.id_comp, category.assoc, - iso.inv_hom_id_app_assoc, eq_to_iso.inv, eq_to_hom_app, eq_to_hom_map, - category_theory.functor.map_comp, functor.image_preimage, - obj_zero_map_μ_app, ε_hom_inv_app_assoc], - have := dcongr_arg (λ a, (i a).inv.app X) (add_zero n), - dsimp at this, - simp [this], - end, } - -end - -/-- When we construct shifts on a subcategory from shifts on the ambient category, -the inclusion functor intertwines the shifts. -/ -@[nolint unused_arguments] -- incorrectly reports that `[full F]` and `[faithful F]` are unused. -def has_shift_of_fully_faithful_comm - (s : A → C ⥤ C) (i : ∀ i, s i ⋙ F ≅ F ⋙ shift_functor D i) (m : A) : - begin - haveI := has_shift_of_fully_faithful F s i, - exact (shift_functor C m) ⋙ F ≅ F ⋙ shift_functor D m - end := -i m - -end category_theory diff --git a/src/category_theory/shift/basic.lean b/src/category_theory/shift/basic.lean new file mode 100644 index 0000000000000..d78bf92733d19 --- /dev/null +++ b/src/category_theory/shift/basic.lean @@ -0,0 +1,733 @@ +/- +Copyright (c) 2020 Scott Morrison. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Scott Morrison, Johan Commelin, Andrew Yang, Joël Riou +-/ +import category_theory.limits.preserves.shapes.zero +import category_theory.monoidal.End +import category_theory.monoidal.discrete + +/-! +# Shift + +A `shift` on a category `C` indexed by a monoid `A` is nothing more than a monoidal functor +from `A` to `C ⥤ C`. A typical example to keep in mind might be the category of +complexes `⋯ → C_{n-1} → C_n → C_{n+1} → ⋯`. It has a shift indexed by `ℤ`, where we assign to +each `n : ℤ` the functor `C ⥤ C` that re-indexes the terms, so the degree `i` term of `shift n C` +would be the degree `i+n`-th term of `C`. + +## Main definitions +* `has_shift`: A typeclass asserting the existence of a shift functor. +* `shift_equiv`: When the indexing monoid is a group, then the functor indexed by `n` and `-n` forms + an self-equivalence of `C`. +* `shift_comm`: When the indexing monoid is commutative, then shifts commute as well. + +## Implementation Notes + +`[has_shift C A]` is implemented using `monoidal_functor (discrete A) (C ⥤ C)`. +However, the API of monodial functors is used only internally: one should use the API of +shifts functors which includes `shift_functor C a : C ⥤ C` for `a : A`, +`shift_functor_zero C A : shift_functor C (0 : A) ≅ 𝟭 C` and +`shift_functor_add C i j : shift_functor C (i + j) ≅ shift_functor C i ⋙ shift_functor C j` +(and its variant `shift_functor_add'`). These isomorphisms satisfy some coherence properties +which are stated in lemmas like `shift_functor_add'_assoc`, `shift_functor_add'_zero_add` and +`shift_functor_add'_add_zero`. + +-/ +namespace category_theory + +noncomputable theory + +universes v u + +variables (C : Type u) (A : Type*) [category.{v} C] + +local attribute [instance] endofunctor_monoidal_category + +section defs + +variables (A C) [add_monoid A] + +/-- A category has a shift indexed by an additive monoid `A` +if there is a monoidal functor from `A` to `C ⥤ C`. -/ +class has_shift (C : Type u) (A : Type*) [category.{v} C] [add_monoid A] := +(shift : monoidal_functor (discrete A) (C ⥤ C)) + +/-- A helper structure to construct the shift functor `(discrete A) ⥤ (C ⥤ C)`. -/ +@[nolint has_nonempty_instance] +structure shift_mk_core := +(F : A → (C ⥤ C)) +(zero : F 0 ≅ 𝟭 C) +(add : Π n m : A, F (n + m) ≅ F n ⋙ F m) +(assoc_hom_app : ∀ (m₁ m₂ m₃ : A) (X : C), + (add (m₁ + m₂) m₃).hom.app X ≫ (F m₃).map ((add m₁ m₂).hom.app X) = + eq_to_hom (by rw [add_assoc]) ≫ (add m₁ (m₂ + m₃)).hom.app X ≫ + (add m₂ m₃).hom.app ((F m₁).obj X)) +(zero_add_hom_app : ∀ (n : A) (X : C), (add 0 n).hom.app X = + eq_to_hom (by dsimp; rw [zero_add]) ≫ (F n).map (zero.inv.app X)) +(add_zero_hom_app : ∀ (n : A) (X : C), (add n 0).hom.app X = + eq_to_hom (by dsimp; rw [add_zero]) ≫ zero.inv.app ((F n).obj X)) + +namespace shift_mk_core + +variables {C A} + +attribute [reassoc] assoc_hom_app + +@[reassoc] +lemma assoc_inv_app (h : shift_mk_core C A) (m₁ m₂ m₃ : A) (X : C) : + (h.F m₃).map ((h.add m₁ m₂).inv.app X) ≫ (h.add (m₁ + m₂) m₃).inv.app X = + (h.add m₂ m₃).inv.app ((h.F m₁).obj X) ≫ (h.add m₁ (m₂ + m₃)).inv.app X ≫ + eq_to_hom (by rw [add_assoc]) := +begin + rw [← cancel_mono ((h.add (m₁ + m₂) m₃).hom.app X ≫ (h.F m₃).map ((h.add m₁ m₂).hom.app X)), + category.assoc, category.assoc, category.assoc, iso.inv_hom_id_app_assoc, ← functor.map_comp, + iso.inv_hom_id_app, functor.map_id, h.assoc_hom_app, eq_to_hom_trans_assoc, eq_to_hom_refl, + category.id_comp, iso.inv_hom_id_app_assoc, iso.inv_hom_id_app], + refl, +end + +lemma zero_add_inv_app (h : shift_mk_core C A) (n : A) (X : C) : + (h.add 0 n).inv.app X = (h.F n).map (h.zero.hom.app X) ≫ + eq_to_hom (by dsimp; rw [zero_add]) := +by rw [← cancel_epi ((h.add 0 n).hom.app X), iso.hom_inv_id_app, h.zero_add_hom_app, + category.assoc, ← functor.map_comp_assoc, iso.inv_hom_id_app, functor.map_id, + category.id_comp, eq_to_hom_trans, eq_to_hom_refl] + +lemma add_zero_inv_app (h : shift_mk_core C A) (n : A) (X : C) : + (h.add n 0).inv.app X = h.zero.hom.app ((h.F n).obj X) ≫ + eq_to_hom (by dsimp; rw [add_zero]) := +by rw [← cancel_epi ((h.add n 0).hom.app X), iso.hom_inv_id_app, h.add_zero_hom_app, + category.assoc, iso.inv_hom_id_app_assoc, eq_to_hom_trans, eq_to_hom_refl] + +end shift_mk_core + +section + +local attribute [simp] eq_to_hom_map +local attribute [reducible] endofunctor_monoidal_category discrete.add_monoidal + +/-- Constructs a `has_shift C A` instance from `shift_mk_core`. -/ +def has_shift_mk (h : shift_mk_core C A) : has_shift C A := +⟨ { ε := h.zero.inv, + μ := λ m n, (h.add m.as n.as).inv, + μ_natural' := by { rintros ⟨X⟩ ⟨Y⟩ ⟨X'⟩ ⟨Y'⟩ ⟨⟨⟨rfl⟩⟩⟩ ⟨⟨⟨rfl⟩⟩⟩, ext, + dsimp, simp only [discrete.functor_map_id, category.assoc], dsimp, simp }, + associativity' := + begin + rintros ⟨m₁⟩ ⟨m₂⟩ ⟨m₃⟩, + ext X, + dsimp, + simp [h.assoc_inv_app_assoc m₁ m₂ m₃ X], + end, + left_unitality' := + begin + rintro ⟨n⟩, + ext X, + dsimp, + simp only [h.zero_add_inv_app, ←functor.map_comp, category.id_comp, eq_to_hom_map, + eq_to_hom_app, category.assoc, eq_to_hom_trans, eq_to_hom_refl, category.comp_id, + iso.inv_hom_id_app], + erw [functor.map_id], + end, + right_unitality' := + begin + rintro ⟨n⟩, + ext X, + dsimp, + simpa only [h.add_zero_inv_app, functor.map_id, category.comp_id, eq_to_hom_map, + eq_to_hom_app, category.assoc, eq_to_hom_trans, eq_to_hom_refl, iso.inv_hom_id_app], + end, + ..(discrete.functor h.F) }⟩ + +end + +section + +variables [has_shift C A] + +/-- The monoidal functor from `A` to `C ⥤ C` given a `has_shift` instance. -/ +def shift_monoidal_functor : monoidal_functor (discrete A) (C ⥤ C) := has_shift.shift + +variable {A} + +/-- The shift autoequivalence, moving objects and morphisms 'up'. -/ +def shift_functor (i : A) : C ⥤ C := (shift_monoidal_functor C A).obj ⟨i⟩ + +/-- Shifting by `i + j` is the same as shifting by `i` and then shifting by `j`. -/ +def shift_functor_add (i j : A) : + shift_functor C (i + j) ≅ shift_functor C i ⋙ shift_functor C j := +((shift_monoidal_functor C A).μ_iso ⟨i⟩ ⟨j⟩).symm + +/-- When `k = i + j`, shifting by `k` is the same as shifting by `i` and then shifting by `j`. -/ +def shift_functor_add' (i j k : A) (h : i + j = k) : + shift_functor C k ≅ shift_functor C i ⋙ shift_functor C j := +eq_to_iso (by rw [h]) ≪≫ shift_functor_add C i j + +lemma shift_functor_add'_eq_shift_functor_add (i j : A) : + shift_functor_add' C i j (i+j) rfl = shift_functor_add C i j := +by { ext1, apply category.id_comp } + +variables (A) + +/-- Shifting by zero is the identity functor. -/ +def shift_functor_zero : shift_functor C (0 : A) ≅ 𝟭 C := +(shift_monoidal_functor C A).ε_iso.symm + +end + +variables {C A} + +lemma shift_mk_core.shift_functor_eq (h : shift_mk_core C A) (a : A) : + @shift_functor _ _ _ _ (has_shift_mk C A h) a = h.F a := +functor.ext (by tidy) (by tidy) + +lemma shift_mk_core.shift_functor_zero_eq (h : shift_mk_core C A) : + @shift_functor_zero _ _ _ _ (has_shift_mk C A h) = h.zero := +begin + letI := has_shift_mk C A h, + ext1, + suffices : (shift_functor_zero C A).inv = h.zero.inv, + { rw [← cancel_mono (h.zero.inv), h.zero.hom_inv_id, ← this, iso.hom_inv_id], + refl, }, + refl, +end + +lemma shift_mk_core.shift_functor_add_eq (h : shift_mk_core C A) (a b : A) : + @shift_functor_add _ _ _ _ (has_shift_mk C A h) a b = h.add a b := +begin + letI := has_shift_mk C A h, + ext1, + suffices : (shift_functor_add C a b).inv = (h.add a b).inv, + { rw [← cancel_mono ((h.add a b).inv), (h.add a b).hom_inv_id, ← this, iso.hom_inv_id], + refl, }, + refl, +end + +-- Any better notational suggestions? +notation X`⟦`n`⟧`:20 := (shift_functor _ n).obj X +notation f`⟦`n`⟧'`:80 := (shift_functor _ n).map f + +variable (C) + +variables [has_shift C A] + +local attribute [reducible] endofunctor_monoidal_category discrete.add_monoidal + +lemma shift_functor_add'_zero_add (a : A) : + shift_functor_add' C 0 a a (zero_add a) = (functor.left_unitor _).symm ≪≫ + iso_whisker_right (shift_functor_zero C A).symm (shift_functor C a) := +begin + ext X, + dsimp [shift_functor_add'], + erw [obj_ε_app], + simpa [eq_to_hom_map], +end + +lemma shift_functor_add'_add_zero (a : A) : + shift_functor_add' C a 0 a (add_zero a) = (functor.right_unitor _).symm ≪≫ + iso_whisker_left (shift_functor C a) (shift_functor_zero C A).symm := +begin + ext X, + dsimp [shift_functor_add'], + erw [ε_app_obj], + simpa [eq_to_hom_map], +end + +lemma shift_functor_add'_assoc (a₁ a₂ a₃ a₁₂ a₂₃ a₁₂₃ : A) + (h₁₂ : a₁ + a₂ = a₁₂) (h₂₃ : a₂ + a₃ = a₂₃) (h₁₂₃ : a₁ + a₂ + a₃ = a₁₂₃) : + shift_functor_add' C a₁₂ a₃ a₁₂₃ (by rw [← h₁₂, h₁₂₃]) ≪≫ + iso_whisker_right (shift_functor_add' C a₁ a₂ a₁₂ h₁₂) _ ≪≫ functor.associator _ _ _ = + shift_functor_add' C a₁ a₂₃ a₁₂₃ (by rw [← h₂₃, ← add_assoc, h₁₂₃]) ≪≫ + iso_whisker_left _ (shift_functor_add' C a₂ a₃ a₂₃ h₂₃) := +begin + substs h₁₂ h₂₃ h₁₂₃, + ext X, + dsimp, + simp only [shift_functor_add'_eq_shift_functor_add, category.comp_id], + dsimp [shift_functor_add', shift_functor_add, shift_functor], + simp [obj_μ_inv_app, eq_to_hom_map], +end + +lemma shift_functor_add_assoc (a₁ a₂ a₃ : A) : + shift_functor_add C (a₁ + a₂) a₃ ≪≫ + iso_whisker_right (shift_functor_add C a₁ a₂) _ ≪≫ functor.associator _ _ _ = + shift_functor_add' C a₁ (a₂ + a₃) _ (add_assoc a₁ a₂ a₃).symm ≪≫ + iso_whisker_left _ (shift_functor_add C a₂ a₃) := +begin + ext X, + simpa [shift_functor_add'_eq_shift_functor_add] + using nat_trans.congr_app (congr_arg iso.hom + (shift_functor_add'_assoc C a₁ a₂ a₃ _ _ _ rfl rfl rfl)) X, +end + +variable {C} + +lemma shift_functor_add'_zero_add_hom_app (a : A) (X : C) : + (shift_functor_add' C 0 a a (zero_add a)).hom.app X = + ((shift_functor_zero C A).inv.app X)⟦a⟧' := +by simpa using nat_trans.congr_app (congr_arg iso.hom (shift_functor_add'_zero_add C a)) X + +lemma shift_functor_add_zero_add_hom_app (a : A) (X : C) : + (shift_functor_add C 0 a).hom.app X = + eq_to_hom (by dsimp; rw [zero_add]) ≫ ((shift_functor_zero C A).inv.app X)⟦a⟧' := +begin + erw [← shift_functor_add'_zero_add_hom_app], + dsimp [shift_functor_add'], + simp, +end + +lemma shift_functor_add'_zero_add_inv_app (a : A) (X : C) : + (shift_functor_add' C 0 a a (zero_add a)).inv.app X = + ((shift_functor_zero C A).hom.app X)⟦a⟧' := +begin + have := nat_trans.congr_app (congr_arg iso.inv (shift_functor_add'_zero_add C a)) X, + simp only [iso.trans_inv, iso_whisker_right_inv, iso.symm_inv, nat_trans.comp_app, + whisker_right_app, functor.left_unitor_hom_app] at this, + dsimp at this, + simpa only [category.comp_id] using this, +end + +lemma shift_functor_add_zero_add_inv_app (a : A) (X : C) : + (shift_functor_add C 0 a).inv.app X = + ((shift_functor_zero C A).hom.app X)⟦a⟧' ≫ eq_to_hom (by dsimp; rw [zero_add]) := +begin + erw [← shift_functor_add'_zero_add_inv_app], + dsimp [shift_functor_add'], + simp, +end + +lemma shift_functor_add'_add_zero_hom_app (a : A) (X : C): + (shift_functor_add' C a 0 a (add_zero a)).hom.app X = + (shift_functor_zero C A).inv.app (X⟦a⟧) := +by simpa using nat_trans.congr_app (congr_arg iso.hom (shift_functor_add'_add_zero C a)) X + +lemma shift_functor_add_add_zero_hom_app (a : A) (X : C): + (shift_functor_add C a 0).hom.app X = + eq_to_hom (by dsimp; rw [add_zero]) ≫ (shift_functor_zero C A).inv.app (X⟦a⟧) := +begin + rw [← shift_functor_add'_add_zero_hom_app], + dsimp [shift_functor_add'], + simp, +end + +lemma shift_functor_add'_add_zero_inv_app (a : A) (X : C): + (shift_functor_add' C a 0 a (add_zero a)).inv.app X = + (shift_functor_zero C A).hom.app (X⟦a⟧) := +begin + have := nat_trans.congr_app (congr_arg iso.inv (shift_functor_add'_add_zero C a)) X, + simp only [iso.trans_inv, iso_whisker_left_inv, iso.symm_inv, nat_trans.comp_app, + whisker_left_app, functor.right_unitor_hom_app] at this, + dsimp at this, + simpa only [category.comp_id] using this, +end + +lemma shift_functor_add_add_zero_inv_app (a : A) (X : C): + (shift_functor_add C a 0).inv.app X = + (shift_functor_zero C A).hom.app (X⟦a⟧) ≫ eq_to_hom (by dsimp; rw [add_zero]) := +begin + rw [← shift_functor_add'_add_zero_inv_app], + dsimp [shift_functor_add'], + simp, +end + +@[reassoc] +lemma shift_functor_add'_assoc_hom_app (a₁ a₂ a₃ a₁₂ a₂₃ a₁₂₃ : A) + (h₁₂ : a₁ + a₂ = a₁₂) (h₂₃ : a₂ + a₃ = a₂₃) (h₁₂₃ : a₁ + a₂ + a₃ = a₁₂₃) (X : C) : + (shift_functor_add' C a₁₂ a₃ a₁₂₃ (by rw [← h₁₂, h₁₂₃])).hom.app X ≫ + ((shift_functor_add' C a₁ a₂ a₁₂ h₁₂).hom.app X)⟦a₃⟧' = + (shift_functor_add' C a₁ a₂₃ a₁₂₃ (by rw [← h₂₃, ← add_assoc, h₁₂₃])).hom.app X ≫ + (shift_functor_add' C a₂ a₃ a₂₃ h₂₃).hom.app (X⟦a₁⟧) := +by simpa using nat_trans.congr_app (congr_arg iso.hom + (shift_functor_add'_assoc C _ _ _ _ _ _ h₁₂ h₂₃ h₁₂₃)) X + +@[reassoc] +lemma shift_functor_add'_assoc_inv_app (a₁ a₂ a₃ a₁₂ a₂₃ a₁₂₃ : A) + (h₁₂ : a₁ + a₂ = a₁₂) (h₂₃ : a₂ + a₃ = a₂₃) (h₁₂₃ : a₁ + a₂ + a₃ = a₁₂₃) (X : C) : + ((shift_functor_add' C a₁ a₂ a₁₂ h₁₂).inv.app X)⟦a₃⟧' ≫ + (shift_functor_add' C a₁₂ a₃ a₁₂₃ (by rw [← h₁₂, h₁₂₃])).inv.app X = + (shift_functor_add' C a₂ a₃ a₂₃ h₂₃).inv.app (X⟦a₁⟧) ≫ + (shift_functor_add' C a₁ a₂₃ a₁₂₃ (by rw [← h₂₃, ← add_assoc, h₁₂₃])).inv.app X := +by simpa using nat_trans.congr_app (congr_arg iso.inv + (shift_functor_add'_assoc C _ _ _ _ _ _ h₁₂ h₂₃ h₁₂₃)) X + +@[reassoc] +lemma shift_functor_add_assoc_hom_app (a₁ a₂ a₃ : A) (X : C) : + (shift_functor_add C (a₁ + a₂) a₃).hom.app X ≫ + ((shift_functor_add C a₁ a₂).hom.app X)⟦a₃⟧' = + (shift_functor_add' C a₁ (a₂ + a₃) (a₁ + a₂ + a₃) (add_assoc _ _ _).symm).hom.app X ≫ + (shift_functor_add C a₂ a₃).hom.app (X⟦a₁⟧) := +by simpa using nat_trans.congr_app (congr_arg iso.hom + (shift_functor_add_assoc C a₁ a₂ a₃)) X + +@[reassoc] +lemma shift_functor_add_assoc_inv_app (a₁ a₂ a₃ : A) (X : C) : + ((shift_functor_add C a₁ a₂).inv.app X)⟦a₃⟧' ≫ + (shift_functor_add C (a₁ + a₂) a₃).inv.app X = + (shift_functor_add C a₂ a₃).inv.app (X⟦a₁⟧) ≫ + (shift_functor_add' C a₁ (a₂ + a₃) (a₁ + a₂ + a₃) (add_assoc _ _ _).symm).inv.app X := +by simpa using nat_trans.congr_app (congr_arg iso.inv + (shift_functor_add_assoc C a₁ a₂ a₃)) X + +end defs + +section add_monoid + +variables {C A} [add_monoid A] [has_shift C A] (X Y : C) (f : X ⟶ Y) + +@[simp] lemma has_shift.shift_obj_obj (n : A) (X : C) : (has_shift.shift.obj ⟨n⟩).obj X = X⟦n⟧ := +rfl + +/-- Shifting by `i + j` is the same as shifting by `i` and then shifting by `j`. -/ +abbreviation shift_add (i j : A) : X⟦i + j⟧ ≅ X⟦i⟧⟦j⟧ := (shift_functor_add C i j).app _ + +lemma shift_shift' (i j : A) : + f⟦i⟧'⟦j⟧' = (shift_add X i j).inv ≫ f⟦i + j⟧' ≫ (shift_add Y i j).hom := +by { symmetry, apply nat_iso.naturality_1 } + +variables (A) + +/-- Shifting by zero is the identity functor. -/ +abbreviation shift_zero : + X⟦0⟧ ≅ X := (shift_functor_zero C A).app _ + +lemma shift_zero' : + f⟦(0 : A)⟧' = (shift_zero A X).hom ≫ f ≫ (shift_zero A Y).inv := +by { symmetry, apply nat_iso.naturality_2 } + +variables (C) {A} + +/-- When `i + j = 0`, shifting by `i` and by `j` gives the identity functor -/ +def shift_functor_comp_iso_id (i j : A) (h : i + j = 0) : + shift_functor C i ⋙ shift_functor C j ≅ 𝟭 C := +(shift_functor_add' C i j 0 h).symm ≪≫ shift_functor_zero C A + +end add_monoid + +section add_group + +variables (C) {A} [add_group A] [has_shift C A] + +/-- shifting by `i` and shifting by `j` forms an equivalence when `i + j = 0`. -/ +@[simps] +def shift_equiv' (i j : A) (h : i + j = 0) : C ≌ C := +{ functor := shift_functor C i, + inverse := shift_functor C j, + unit_iso := (shift_functor_comp_iso_id C i j h).symm, + counit_iso := (shift_functor_comp_iso_id C j i + (by rw [← add_left_inj j, add_assoc, h, zero_add, add_zero])), + functor_unit_iso_comp' := λ X, begin + let E := equiv_of_tensor_iso_unit (shift_monoidal_functor C A) ⟨i⟩ ⟨j⟩ + (eq_to_iso (by ext; exact h)) + (eq_to_iso (by ext; dsimp; rw [← add_left_inj j, add_assoc, h, zero_add, add_zero])) + (subsingleton.elim _ _), + convert equivalence.functor_unit_iso_comp E X, + all_goals + { ext X, + dsimp [shift_functor_comp_iso_id, unit_of_tensor_iso_unit, shift_functor_add'], + simpa only [eq_to_hom_map, category.assoc], }, + end } + +/-- shifting by `n` and shifting by `-n` forms an equivalence. -/ +abbreviation shift_equiv (i : A) : C ≌ C := shift_equiv' C i (-i) (add_neg_self i) + +variables (X Y : C) (f : X ⟶ Y) + +/-- Shifting by `i` is an equivalence. -/ +instance (i : A) : is_equivalence (shift_functor C i) := +is_equivalence.of_equivalence (shift_equiv C i) + +@[simp] lemma shift_functor_inv (i : A) : + (shift_functor C i).inv = shift_functor C (-i) := +rfl + +section + +variables (C) + +/-- Shifting by `n` is an essentially surjective functor. -/ +instance shift_functor_ess_surj (i : A) : ess_surj (shift_functor C i) := + equivalence.ess_surj_of_equivalence _ + +end + +variables {C} + +/-- Shifting by `i` and then shifting by `-i` is the identity. -/ +abbreviation shift_shift_neg (i : A) : X⟦i⟧⟦-i⟧ ≅ X := +(shift_equiv C i).unit_iso.symm.app _ + +/-- Shifting by `-i` and then shifting by `i` is the identity. -/ +abbreviation shift_neg_shift (i : A) : X⟦-i⟧⟦i⟧ ≅ X := +(shift_equiv C i).counit_iso.app _ + +variables {X Y} + +lemma shift_shift_neg' (i : A) : + f⟦i⟧'⟦-i⟧' = (shift_functor_comp_iso_id C i (-i) (add_neg_self i)).hom.app X ≫ + f ≫ (shift_functor_comp_iso_id C i (-i) (add_neg_self i)).inv.app Y := +(nat_iso.naturality_2 (shift_functor_comp_iso_id C i (-i) (add_neg_self i)) f).symm + +lemma shift_neg_shift' (i : A) : + f⟦-i⟧'⟦i⟧' = (shift_functor_comp_iso_id C (-i) i (neg_add_self i)).hom.app X ≫ f ≫ + (shift_functor_comp_iso_id C (-i) i (neg_add_self i)).inv.app Y := +(nat_iso.naturality_2 (shift_functor_comp_iso_id C (-i) i (neg_add_self i)) f).symm + +lemma shift_equiv_triangle (n : A) (X : C) : + (shift_shift_neg X n).inv⟦n⟧' ≫ (shift_neg_shift (X⟦n⟧) n).hom = 𝟙 (X⟦n⟧) := +(shift_equiv C n).functor_unit_iso_comp X + +section + +lemma shift_shift_functor_comp_iso_id_hom_app (n m : A) (h : n + m = 0) (X : C) : + ((shift_functor_comp_iso_id C n m h).hom.app X)⟦n⟧' = + (shift_functor_comp_iso_id C m n + (by rw [← neg_eq_of_add_eq_zero_left h, add_right_neg])).hom.app (X⟦n⟧) := +begin + dsimp [shift_functor_comp_iso_id], + simpa only [functor.map_comp, ← shift_functor_add'_zero_add_inv_app n X, + ← shift_functor_add'_add_zero_inv_app n X ] + using shift_functor_add'_assoc_inv_app n m n 0 0 n h + (by rw [← neg_eq_of_add_eq_zero_left h, add_right_neg]) (by rw [h, zero_add]) X, +end + +lemma shift_shift_functor_comp_iso_id_inv_app (n m : A) (h : n + m = 0) (X : C) : + ((shift_functor_comp_iso_id C n m h).inv.app X)⟦n⟧' = + ((shift_functor_comp_iso_id C m n + (by rw [← neg_eq_of_add_eq_zero_left h, add_right_neg])).inv.app (X⟦n⟧)) := +begin + rw [← cancel_mono (((shift_functor_comp_iso_id C n m h).hom.app X)⟦n⟧'), + ← functor.map_comp, iso.inv_hom_id_app, functor.map_id, + shift_shift_functor_comp_iso_id_hom_app, iso.inv_hom_id_app], + refl, +end + +lemma shift_shift_functor_comp_iso_id_add_neg_self_hom_app (n : A) (X : C) : + ((shift_functor_comp_iso_id C n (-n) (add_neg_self n)).hom.app X)⟦n⟧' = + (shift_functor_comp_iso_id C (-n) n (neg_add_self n)).hom.app (X⟦n⟧) := +by apply shift_shift_functor_comp_iso_id_hom_app + +lemma shift_shift_functor_comp_iso_id_add_neg_self_inv_app (n : A) (X : C) : + ((shift_functor_comp_iso_id C n (-n) (add_neg_self n)).inv.app X)⟦n⟧' = + (shift_functor_comp_iso_id C (-n) n (neg_add_self n)).inv.app (X⟦n⟧) := +by apply shift_shift_functor_comp_iso_id_inv_app + +lemma shift_shift_functor_comp_iso_id_neg_add_self_hom_app (n : A) (X : C) : + ((shift_functor_comp_iso_id C (-n) n (neg_add_self n)).hom.app X)⟦-n⟧' = + (shift_functor_comp_iso_id C n (-n) (add_neg_self n)).hom.app (X⟦-n⟧) := +by apply shift_shift_functor_comp_iso_id_hom_app + +lemma shift_shift_functor_comp_iso_id_neg_add_self_inv_app (n : A) (X : C) : + ((shift_functor_comp_iso_id C (-n) n (neg_add_self n)).inv.app X)⟦-n⟧' = + (shift_functor_comp_iso_id C n (-n) (add_neg_self n)).inv.app (X⟦-n⟧) := +by apply shift_shift_functor_comp_iso_id_inv_app + +end + +variables {A C} + +open category_theory.limits + +variables [has_zero_morphisms C] + +lemma shift_zero_eq_zero (X Y : C) (n : A) : (0 : X ⟶ Y)⟦n⟧' = (0 : X⟦n⟧ ⟶ Y⟦n⟧) := +category_theory.functor.map_zero _ _ _ + +end add_group + +section add_comm_monoid + +variables (C) {A} [add_comm_monoid A] [has_shift C A] + +/-- When shifts are indexed by an additive commutative monoid, then shifts commute. -/ +def shift_functor_comm (i j : A) : + shift_functor C i ⋙ shift_functor C j ≅ + shift_functor C j ⋙ shift_functor C i := +(shift_functor_add C i j).symm ≪≫ shift_functor_add' C j i (i + j) (add_comm j i) + +lemma shift_functor_comm_eq (i j k : A) (h : i + j = k): + shift_functor_comm C i j = (shift_functor_add' C i j k h).symm ≪≫ + shift_functor_add' C j i k (by rw [add_comm j i, h]) := +begin + subst h, + rw [shift_functor_add'_eq_shift_functor_add], + refl, +end + +lemma shift_functor_comm_symm (i j : A) : + (shift_functor_comm C i j).symm = shift_functor_comm C j i := +begin + ext1, + dsimp, + simpa only [shift_functor_comm_eq C i j (i + j) rfl, + shift_functor_comm_eq C j i (i + j) (add_comm j i)], +end + +variables {C} (X Y : C) (f : X ⟶ Y) + +/-- When shifts are indexed by an additive commutative monoid, then shifts commute. -/ +abbreviation shift_comm (i j : A) : X⟦i⟧⟦j⟧ ≅ X⟦j⟧⟦i⟧ := + (shift_functor_comm C i j).app X + +@[simp] lemma shift_comm_symm (i j : A) : (shift_comm X i j).symm = shift_comm X j i := +begin + ext1, + exact nat_trans.congr_app (congr_arg iso.hom (shift_functor_comm_symm C i j)) X, +end + +variables {X Y} + +/-- When shifts are indexed by an additive commutative monoid, then shifts commute. -/ +lemma shift_comm' (i j : A) : + f⟦i⟧'⟦j⟧' = (shift_comm _ _ _).hom ≫ f⟦j⟧'⟦i⟧' ≫ (shift_comm _ _ _).hom := +begin + erw [← shift_comm_symm Y i j, ← ((shift_functor_comm C i j).hom.naturality_assoc f), + iso.hom_inv_id_app, category.comp_id], + refl, +end + +@[reassoc] lemma shift_comm_hom_comp (i j : A) : + (shift_comm X i j).hom ≫ f⟦j⟧'⟦i⟧' = f⟦i⟧'⟦j⟧' ≫ (shift_comm Y i j).hom := +by rw [shift_comm', ← shift_comm_symm, iso.symm_hom, iso.inv_hom_id_assoc] + +lemma shift_functor_zero_hom_app_shift (n : A) : + (shift_functor_zero C A).hom.app (X⟦n⟧) = + (shift_functor_comm C n 0).hom.app X ≫ ((shift_functor_zero C A).hom.app X)⟦n⟧' := +begin + rw [← shift_functor_add'_zero_add_inv_app n X, shift_functor_comm_eq C n 0 n (add_zero n)], + dsimp, + rw [category.assoc, iso.hom_inv_id_app, category.comp_id, shift_functor_add'_add_zero_inv_app], +end + +lemma shift_functor_zero_inv_app_shift (n : A) : + (shift_functor_zero C A).inv.app (X⟦n⟧) = + ((shift_functor_zero C A).inv.app X)⟦n⟧' ≫ (shift_functor_comm C n 0).inv.app X := +begin + rw [← cancel_mono ((shift_functor_zero C A).hom.app (X⟦n⟧)), category.assoc, iso.inv_hom_id_app, + shift_functor_zero_hom_app_shift, iso.inv_hom_id_app_assoc, ← functor.map_comp, + iso.inv_hom_id_app], + dsimp, + rw [functor.map_id], +end + +@[reassoc] +lemma shift_functor_comm_hom_app_comp_shift_shift_functor_add_hom_app (m₁ m₂ m₃ : A) (X : C) : + (shift_functor_comm C m₁ (m₂ + m₃)).hom.app X ≫ + ((shift_functor_add C m₂ m₃).hom.app X)⟦m₁⟧' = + (shift_functor_add C m₂ m₃).hom.app (X⟦m₁⟧) ≫ + ((shift_functor_comm C m₁ m₂).hom.app X)⟦m₃⟧' ≫ + (shift_functor_comm C m₁ m₃).hom.app (X⟦m₂⟧) := +begin + simp only [← cancel_mono ((shift_functor_comm C m₁ m₃).inv.app (X⟦m₂⟧)), + ← cancel_mono (((shift_functor_comm C m₁ m₂).inv.app X)⟦m₃⟧'), + category.assoc, iso.hom_inv_id_app], + dsimp, + simp only [category.id_comp, ← functor.map_comp, iso.hom_inv_id_app], + dsimp, + simp only [functor.map_id, category.comp_id, + shift_functor_comm_eq C _ _ _ rfl, ← shift_functor_add'_eq_shift_functor_add], + dsimp, + simp only [category.assoc, iso.hom_inv_id_app_assoc, iso.inv_hom_id_app_assoc, + ← functor.map_comp, + shift_functor_add'_assoc_hom_app_assoc m₂ m₃ m₁ (m₂ + m₃) (m₁ + m₃) (m₁ + (m₂ + m₃)) rfl + (add_comm m₃ m₁) (add_comm _ m₁) X, + ← shift_functor_add'_assoc_hom_app_assoc m₂ m₁ m₃ (m₁ + m₂) (m₁ + m₃) + (m₁ + (m₂ + m₃)) (add_comm _ _) rfl (by rw [add_comm m₂ m₁, add_assoc]) X, + shift_functor_add'_assoc_hom_app m₁ m₂ m₃ + (m₁ + m₂) (m₂ + m₃) (m₁ + (m₂ + m₃)) rfl rfl (add_assoc _ _ _) X], +end + +end add_comm_monoid + +variables {C A} {D : Type*} [category D] [add_monoid A] [has_shift D A] +variables (F : C ⥤ D) [full F] [faithful F] + +section + +variables (s : A → C ⥤ C) (i : ∀ i, s i ⋙ F ≅ F ⋙ shift_functor D i) + +include F s i + +/-- auxiliary definition for `has_shift_of_fully_faithful` -/ +def has_shift_of_fully_faithful_zero : s 0 ≅ 𝟭 C := +nat_iso_of_comp_fully_faithful F ((i 0) ≪≫ iso_whisker_left F (shift_functor_zero D A) ≪≫ + functor.right_unitor _ ≪≫ (functor.left_unitor _).symm) + +@[simp] +lemma map_has_shift_of_fully_faithful_zero_hom_app (X : C) : + F.map ((has_shift_of_fully_faithful_zero F s i).hom.app X) = + (i 0).hom.app X ≫ (shift_functor_zero D A).hom.app (F.obj X) := +by { dsimp [has_shift_of_fully_faithful_zero], simp, } + +@[simp] +lemma map_has_shift_of_fully_faithful_zero_inv_app (X : C) : +F.map ((has_shift_of_fully_faithful_zero F s i).inv.app X) = + (shift_functor_zero D A).inv.app (F.obj X) ≫ (i 0).inv.app X := +by { dsimp [has_shift_of_fully_faithful_zero], simp, } + +/-- auxiliary definition for `has_shift_of_fully_faithful` -/ +def has_shift_of_fully_faithful_add (a b : A) : s (a + b) ≅ s a ⋙ s b := +nat_iso_of_comp_fully_faithful F (i (a + b) ≪≫ + iso_whisker_left _ (shift_functor_add D a b) ≪≫ + (functor.associator _ _ _).symm ≪≫ (iso_whisker_right (i a).symm _) ≪≫ + functor.associator _ _ _ ≪≫ (iso_whisker_left _ (i b).symm) ≪≫ + (functor.associator _ _ _).symm) + +@[simp] +lemma map_has_shift_of_fully_faithful_add_hom_app (a b : A) (X : C) : + F.map ((has_shift_of_fully_faithful_add F s i a b).hom.app X) = + (i (a + b)).hom.app X ≫ (shift_functor_add D a b).hom.app (F.obj X) ≫ + ((i a).inv.app X)⟦b⟧' ≫ (i b).inv.app ((s a).obj X) := +by { dsimp [has_shift_of_fully_faithful_add], simp, } + +@[simp] +lemma map_has_shift_of_fully_faithful_add_inv_app (a b : A) (X : C) : + F.map ((has_shift_of_fully_faithful_add F s i a b).inv.app X) = + (i b).hom.app ((s a).obj X) ≫ ((i a).hom.app X)⟦b⟧' ≫ + (shift_functor_add D a b).inv.app (F.obj X) ≫ (i (a + b)).inv.app X := +by { dsimp [has_shift_of_fully_faithful_add], simp, } + +/-- Given a family of endomorphisms of `C` which are interwined by a fully faithful `F : C ⥤ D` +with shift functors on `D`, we can promote that family to shift functors on `C`. -/ +def has_shift_of_fully_faithful : has_shift C A := has_shift_mk C A + { F := s, + zero := has_shift_of_fully_faithful_zero F s i, + add := has_shift_of_fully_faithful_add F s i, + assoc_hom_app := λ m₁ m₂ m₃ X, F.map_injective begin + rw [← cancel_mono ((i m₃).hom.app ((s m₂).obj ((s m₁).obj X)))], + simp only [functor.map_comp, map_has_shift_of_fully_faithful_add_hom_app, category.assoc, + iso.inv_hom_id_app_assoc, iso.inv_hom_id_app], + erw [(i m₃).hom.naturality], + have := dcongr_arg (λ a, (i a).hom.app X) (add_assoc m₁ m₂ m₃), + dsimp at this, + dsimp, + rw [iso.inv_hom_id_app_assoc, map_has_shift_of_fully_faithful_add_hom_app, this, + eq_to_hom_map, category.comp_id, ← functor.map_comp, category.assoc, + iso.inv_hom_id_app_assoc, functor.map_comp, functor.map_comp, + shift_functor_add_assoc_hom_app_assoc m₁ m₂ m₃ (F.obj X)], + dsimp [shift_functor_add'], + simp only [eq_to_hom_app, category.assoc, eq_to_hom_trans_assoc, eq_to_hom_refl, + category.id_comp, nat_trans.naturality_assoc, functor.comp_map], + end, + zero_add_hom_app := λ n X, F.map_injective begin + have this := dcongr_arg (λ a, (i a).hom.app X) (zero_add n), + dsimp at this, + rw [← cancel_mono ((i n).hom.app ((s 0).obj X))], + simp only [this, map_has_shift_of_fully_faithful_add_hom_app, + shift_functor_add_zero_add_hom_app, eq_to_hom_map, category.assoc, + eq_to_hom_trans_assoc, eq_to_hom_refl, category.id_comp, iso.inv_hom_id_app, + functor.map_comp], + erw [(i n).hom.naturality], + dsimp, + simp, + end, + add_zero_hom_app := λ n X, F.map_injective begin + have := dcongr_arg (λ a, (i a).hom.app X) (add_zero n), + dsimp at this, + simpa [this, ← nat_trans.naturality_assoc, eq_to_hom_map, + shift_functor_add_add_zero_hom_app], + end, } + +end + +end category_theory diff --git a/src/category_theory/triangulated/basic.lean b/src/category_theory/triangulated/basic.lean index 7cfece83be0f9..8b1ab1d9950bf 100644 --- a/src/category_theory/triangulated/basic.lean +++ b/src/category_theory/triangulated/basic.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Luke Kershaw -/ import data.int.basic -import category_theory.shift +import category_theory.shift.basic /-! # Triangles @@ -135,4 +135,33 @@ instance triangle_category : category (triangle C) := id := λ A, triangle_morphism_id A, comp := λ A B C f g, f.comp g } +/-- a constructor for morphisms of triangles -/ +@[simps] +def triangle.hom_mk (A B : triangle C) + (hom₁ : A.obj₁ ⟶ B.obj₁) (hom₂ : A.obj₂ ⟶ B.obj₂) (hom₃ : A.obj₃ ⟶ B.obj₃) + (comm₁ : A.mor₁ ≫ hom₂ = hom₁ ≫ B.mor₁) (comm₂ : A.mor₂ ≫ hom₃ = hom₂ ≫ B.mor₂) + (comm₃ : A.mor₃ ≫ hom₁⟦1⟧' = hom₃ ≫ B.mor₃) : A ⟶ B := +{ hom₁ := hom₁, + hom₂ := hom₂, + hom₃ := hom₃, + comm₁' := comm₁, + comm₂' := comm₂, + comm₃' := comm₃, } + +/-- a constructor for isomorphisms of triangles -/ +@[simps] +def triangle.iso_mk (A B : triangle C) + (iso₁ : A.obj₁ ≅ B.obj₁) (iso₂ : A.obj₂ ≅ B.obj₂) (iso₃ : A.obj₃ ≅ B.obj₃) + (comm₁ : A.mor₁ ≫ iso₂.hom = iso₁.hom ≫ B.mor₁) + (comm₂ : A.mor₂ ≫ iso₃.hom = iso₂.hom ≫ B.mor₂) + (comm₃ : A.mor₃ ≫ iso₁.hom⟦1⟧' = iso₃.hom ≫ B.mor₃) : A ≅ B := +{ hom := triangle.hom_mk _ _ iso₁.hom iso₂.hom iso₃.hom comm₁ comm₂ comm₃, + inv := triangle.hom_mk _ _ iso₁.inv iso₂.inv iso₃.inv + (by simp only [← cancel_mono iso₂.hom, assoc, iso.inv_hom_id, comp_id, + comm₁, iso.inv_hom_id_assoc]) + (by simp only [← cancel_mono iso₃.hom, assoc, iso.inv_hom_id, comp_id, + comm₂, iso.inv_hom_id_assoc]) + (by simp only [← cancel_mono (iso₁.hom⟦(1 : ℤ)⟧'), assoc, ← functor.map_comp, + iso.inv_hom_id, category_theory.functor.map_id, comp_id, comm₃, iso.inv_hom_id_assoc]), } + end category_theory.pretriangulated diff --git a/src/category_theory/triangulated/pretriangulated.lean b/src/category_theory/triangulated/pretriangulated.lean index 458d221180f72..4f101a20de6a3 100644 --- a/src/category_theory/triangulated/pretriangulated.lean +++ b/src/category_theory/triangulated/pretriangulated.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Luke Kershaw -/ import category_theory.preadditive.additive_functor -import category_theory.shift +import category_theory.shift.basic import category_theory.triangulated.rotate /-! @@ -143,87 +143,5 @@ TODO: If `C` is pretriangulated with respect to a shift, then `Cᵒᵖ` is pretriangulated with respect to the inverse shift. -/ -omit hC - -/-- -The underlying structure of a triangulated functor between pretriangulated categories `C` and `D` -is a functor `F : C ⥤ D` together with given functorial isomorphisms `ξ X : F(X⟦1⟧) ⟶ F(X)⟦1⟧`. --/ -structure triangulated_functor_struct extends (C ⥤ D) := -(comm_shift : shift_functor C (1 : ℤ) ⋙ to_functor ≅ to_functor ⋙ shift_functor D (1 : ℤ)) - -namespace triangulated_functor_struct - -/-- The identity `triangulated_functor_struct`. -/ -def id : triangulated_functor_struct C C := -{ obj := λ X, X, - map := λ _ _ f, f, - comm_shift := by refl } - -instance : inhabited (triangulated_functor_struct C C) := ⟨id C⟩ - -variables {C D} -/-- -Given a `triangulated_functor_struct` we can define a functor from triangles of `C` to -triangles of `D`. --/ -@[simps] -def map_triangle (F : triangulated_functor_struct C D) : triangle C ⥤ triangle D := -{ obj := λ T, triangle.mk (F.map T.mor₁) (F.map T.mor₂) - (F.map T.mor₃ ≫ F.comm_shift.hom.app T.obj₁), - map := λ S T f, - { hom₁ := F.map f.hom₁, - hom₂ := F.map f.hom₂, - hom₃ := F.map f.hom₃, - comm₁' := by { dsimp, simp only [←F.to_functor.map_comp, f.comm₁], }, - comm₂' := by { dsimp, simp only [←F.to_functor.map_comp, f.comm₂], }, - comm₃' := begin - dsimp, - erw [category.assoc, ←F.comm_shift.hom.naturality], - simp only [functor.comp_map, ←F.to_functor.map_comp_assoc, f.comm₃], - end, }, } - -end triangulated_functor_struct - -include hC -variables (C D) [pretriangulated D] - -/-- -A triangulated functor between pretriangulated categories `C` and `D` is a functor `F : C ⥤ D` -together with given functorial isomorphisms `ξ X : F(X⟦1⟧) ⟶ F(X)⟦1⟧` such that for every -distinguished triangle `(X,Y,Z,f,g,h)` of `C`, the triangle -`(F(X), F(Y), F(Z), F(f), F(g), F(h) ≫ (ξ X))` is a distinguished triangle of `D`. -See --/ -structure triangulated_functor extends triangulated_functor_struct C D := -(map_distinguished' : Π (T : triangle C), (T ∈ dist_triang C) → - (to_triangulated_functor_struct.map_triangle.obj T ∈ dist_triang D) ) - -instance : inhabited (triangulated_functor C C) := -⟨{obj := λ X, X, - map := λ _ _ f, f, - comm_shift := by refl , - map_distinguished' := begin - rintros ⟨_,_,_,_⟩ Tdt, - dsimp at *, - rwa category.comp_id, - end }⟩ - -variables {C D} -/-- -Given a `triangulated_functor` we can define a functor from triangles of `C` to triangles of `D`. --/ -@[simps] -def triangulated_functor.map_triangle (F : triangulated_functor C D) : - triangle C ⥤ triangle D := -F.to_triangulated_functor_struct.map_triangle - -/-- -Given a `triangulated_functor` and a distinguished triangle `T` of `C`, then the triangle it -maps onto in `D` is also distinguished. --/ -lemma triangulated_functor.map_distinguished (F : triangulated_functor C D) (T : triangle C) - (h : T ∈ dist_triang C) : (F.map_triangle.obj T) ∈ dist_triang D := F.map_distinguished' T h - end pretriangulated end category_theory diff --git a/src/category_theory/triangulated/rotate.lean b/src/category_theory/triangulated/rotate.lean index 4e5f57120189f..ebc422b3e55e8 100644 --- a/src/category_theory/triangulated/rotate.lean +++ b/src/category_theory/triangulated/rotate.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2021 Luke Kershaw. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Luke Kershaw +Authors: Luke Kershaw, Joël Riou -/ import category_theory.preadditive.additive_functor import category_theory.triangulated.basic @@ -50,7 +50,6 @@ applying `rotate` gives a triangle of the form: def triangle.rotate (T : triangle C) : triangle C := triangle.mk T.mor₂ T.mor₃ (-T.mor₁⟦1⟧') section -local attribute [semireducible] shift_shift_neg shift_neg_shift /-- Given a triangle of the form: @@ -73,92 +72,6 @@ triangle.mk (-T.mor₃⟦(-1:ℤ)⟧' ≫ (shift_shift_neg _ _).hom) T.mor₁ end -namespace triangle_morphism -variables {T₁ T₂ T₃ T₄: triangle C} -open triangle -/-- -You can also rotate a triangle morphism to get a morphism between the two rotated triangles. -Given a triangle morphism of the form: -``` - f g h - X ───> Y ───> Z ───> X⟦1⟧ - │ │ │ │ - │a │b │c │a⟦1⟧ - V V V V - X' ───> Y' ───> Z' ───> X'⟦1⟧ - f' g' h' -``` -applying `rotate` gives a triangle morphism of the form: - -``` - g h -f⟦1⟧ - Y ───> Z ───> X⟦1⟧ ───> Y⟦1⟧ - │ │ │ │ - │b │c │a⟦1⟧ │b⟦1⟧' - V V V V - Y' ───> Z' ───> X'⟦1⟧ ───> Y'⟦1⟧ - g' h' -f'⟦1⟧ -``` --/ -@[simps] -def rotate (f : triangle_morphism T₁ T₂) : - triangle_morphism (T₁.rotate) (T₂.rotate):= -{ hom₁ := f.hom₂, - hom₂ := f.hom₃, - hom₃ := f.hom₁⟦1⟧', - comm₃' := begin - dsimp, - simp only [rotate_mor₃, comp_neg, neg_comp, ← functor.map_comp, f.comm₁] - end } - -/-- -Given a triangle morphism of the form: -``` - f g h - X ───> Y ───> Z ───> X⟦1⟧ - │ │ │ │ - │a │b │c │a⟦1⟧ - V V V V - X' ───> Y' ───> Z' ───> X'⟦1⟧ - f' g' h' -``` -applying `inv_rotate` gives a triangle morphism that can be thought of as: -``` - -h⟦-1⟧ f g - Z⟦-1⟧ ───> X ───> Y ───> Z - │ │ │ │ - │c⟦-1⟧' │a │b │c - V V V V - Z'⟦-1⟧ ───> X' ───> Y' ───> Z' - -h'⟦-1⟧ f' g' -``` -(note that this diagram doesn't technically fit the definition of triangle morphism, -as `Z⟦-1⟧⟦1⟧` is not necessarily equal to `Z`, and `Z'⟦-1⟧⟦1⟧` is not necessarily equal to `Z'`, -but they are isomorphic, by the `counit_iso` of `shift C`) --/ -@[simps] -def inv_rotate (f : triangle_morphism T₁ T₂) : - triangle_morphism (T₁.inv_rotate) (T₂.inv_rotate) := -{ hom₁ := f.hom₃⟦-1⟧', - hom₂ := f.hom₁, - hom₃ := f.hom₂, - comm₁' := begin - dsimp [inv_rotate_mor₁], - simp only [discrete.functor_map_id, id_comp, preadditive.comp_neg, assoc, - neg_inj, nat_trans.id_app, preadditive.neg_comp], - rw [← functor.map_comp_assoc, ← f.comm₃, functor.map_comp_assoc, μ_naturality_assoc, - nat_trans.naturality, functor.id_map], - end, - comm₃' := begin - dsimp, - simp only [discrete.functor_map_id, id_comp, μ_inv_naturality, - category.assoc, nat_trans.id_app, unit_of_tensor_iso_unit_inv_app], - erw ε_naturality_assoc, - rw comm₂_assoc - end } - -end triangle_morphism - variables (C) /-- @@ -167,7 +80,13 @@ Rotating triangles gives an endofunctor on the category of triangles in `C`. @[simps] def rotate : triangle C ⥤ triangle C := { obj := triangle.rotate, - map := λ _ _ f, f.rotate } + map := λ T₁ T₂ f, + { hom₁ := f.hom₂, + hom₂ := f.hom₃, + hom₃ := f.hom₁⟦1⟧', + comm₃' := by { dsimp, simp only [comp_neg, neg_comp, ← functor.map_comp, f.comm₁], }, }, } + +example : ℕ:= 42 /-- The inverse rotation of triangles gives an endofunctor on the category of triangles in `C`. @@ -175,152 +94,43 @@ The inverse rotation of triangles gives an endofunctor on the category of triang @[simps] def inv_rotate : triangle C ⥤ triangle C := { obj := triangle.inv_rotate, - map := λ _ _ f, f.inv_rotate } + map := λ T₁ T₂ f, + { hom₁ := f.hom₃⟦-1⟧', + hom₂ := f.hom₁, + hom₃ := f.hom₂, + comm₁' := + begin + dsimp, + rw [neg_comp, assoc, comp_neg, neg_inj, ← functor.map_comp_assoc, ← f.comm₃, + functor.map_comp, assoc], + erw [← nat_trans.naturality], + refl, + end, + comm₃' := by { dsimp, erw [← f.comm₂_assoc, assoc, ← nat_trans.naturality], refl, }, }, } variables {C} variables [∀ n : ℤ, functor.additive (shift_functor C n)] -/-- There is a natural map from a triangle to the `inv_rotate` of its `rotate`. -/ -@[simps] -def to_inv_rotate_rotate (T : triangle C) : T ⟶ (inv_rotate C).obj ((rotate C).obj T) := -{ hom₁ := (shift_shift_neg _ _).inv, - hom₂ := 𝟙 T.obj₂, - hom₃ := 𝟙 T.obj₃, - comm₃' := begin - dsimp, - simp only [ε_app_obj, eq_to_iso.hom, discrete.functor_map_id, id_comp, eq_to_iso.inv, - category.assoc, obj_μ_inv_app, functor.map_comp, nat_trans.id_app, obj_ε_app, - unit_of_tensor_iso_unit_inv_app], - erw μ_inv_hom_app_assoc, - refl - end } - -/-- -There is a natural transformation between the identity functor on triangles in `C`, -and the composition of a rotation with an inverse rotation. --/ -@[simps] -def rot_comp_inv_rot_hom : 𝟭 (triangle C) ⟶ rotate C ⋙ inv_rotate C := -{ app := to_inv_rotate_rotate, - naturality' := begin - introv, ext, - { dsimp, - simp only [nat_iso.cancel_nat_iso_inv_right_assoc, discrete.functor_map_id, id_comp, - μ_inv_naturality, assoc, nat_trans.id_app, unit_of_tensor_iso_unit_inv_app], - erw ε_naturality }, - { dsimp, rw [comp_id, id_comp] }, - { dsimp, rw [comp_id, id_comp] }, - end } - -/-- There is a natural map from the `inv_rotate` of the `rotate` of a triangle to itself. -/ -@[simps] -def from_inv_rotate_rotate (T : triangle C) : (inv_rotate C).obj ((rotate C).obj T) ⟶ T := -{ hom₁ := (shift_equiv C 1).unit_inv.app T.obj₁, - hom₂ := 𝟙 T.obj₂, - hom₃ := 𝟙 T.obj₃, - comm₃' := begin - dsimp, - rw [unit_of_tensor_iso_unit_inv_app, ε_app_obj], - simp only [discrete.functor_map_id, nat_trans.id_app, id_comp, assoc, functor.map_comp, - obj_μ_app, obj_ε_inv_app, comp_id, μ_inv_hom_app_assoc], - erw [μ_inv_hom_app, μ_inv_hom_app_assoc, category.comp_id] - end } +local attribute [simp] shift_shift_neg' shift_neg_shift' + shift_shift_functor_comp_iso_id_add_neg_self_inv_app + shift_shift_functor_comp_iso_id_add_neg_self_hom_app -/-- -There is a natural transformation between the composition of a rotation with an inverse rotation -on triangles in `C`, and the identity functor. --/ -@[simps] -def rot_comp_inv_rot_inv : rotate C ⋙ inv_rotate C ⟶ 𝟭 (triangle C) := -{ app := from_inv_rotate_rotate } - -/-- -The natural transformations between the identity functor on triangles in `C` and the composition -of a rotation with an inverse rotation are natural isomorphisms (they are isomorphisms in the -category of functors). --/ +/-- The unit isomorphism of the auto-equivalence of categories `triangle_rotation C` of +`triangle C` given by the rotation of triangles. -/ @[simps] def rot_comp_inv_rot : 𝟭 (triangle C) ≅ rotate C ⋙ inv_rotate C := -{ hom := rot_comp_inv_rot_hom, - inv := rot_comp_inv_rot_inv } - -/-- There is a natural map from the `rotate` of the `inv_rotate` of a triangle to itself. -/ -@[simps] -def from_rotate_inv_rotate (T : triangle C) : (rotate C).obj ((inv_rotate C).obj T) ⟶ T := -{ hom₁ := 𝟙 T.obj₁, - hom₂ := 𝟙 T.obj₂, - hom₃ := (shift_equiv C 1).counit.app T.obj₃, - comm₂' := begin - dsimp, - rw unit_of_tensor_iso_unit_inv_app, - simp only [discrete.functor_map_id, nat_trans.id_app, - id_comp, add_neg_equiv_counit_iso_hom, eq_to_hom_refl, nat_trans.comp_app, assoc, - μ_inv_hom_app_assoc, ε_hom_inv_app], - exact category.comp_id _, - end, - comm₃' := begin - dsimp, - simp only [discrete.functor_map_id, nat_trans.id_app, id_comp, functor.map_neg, - functor.map_comp, obj_μ_app, obj_ε_inv_app, comp_id, assoc, μ_naturality_assoc, neg_neg, - category_theory.functor.map_id, add_neg_equiv_counit_iso_hom, eq_to_hom_refl, - nat_trans.comp_app], - erw [μ_inv_hom_app, category.comp_id, obj_zero_map_μ_app], - rw [discrete.functor_map_id, nat_trans.id_app, comp_id], - end } - -/-- -There is a natural transformation between the composition of an inverse rotation with a rotation -on triangles in `C`, and the identity functor. --/ -@[simps] -def inv_rot_comp_rot_hom : inv_rotate C ⋙ rotate C ⟶ 𝟭 (triangle C) := -{ app := from_rotate_inv_rotate } +nat_iso.of_components (λ T, triangle.iso_mk _ _ ((shift_equiv C (1 : ℤ)).unit_iso.app T.obj₁) + (iso.refl _) (iso.refl _) (by tidy) (by tidy) (by tidy)) (by tidy) -/-- There is a natural map from a triangle to the `rotate` of its `inv_rotate`. -/ -@[simps] -def to_rotate_inv_rotate (T : triangle C) : T ⟶ (rotate C).obj ((inv_rotate C).obj T) := -{ hom₁ := 𝟙 T.obj₁, - hom₂ := 𝟙 T.obj₂, - hom₃ := (shift_equiv C 1).counit_inv.app T.obj₃, - comm₃' := begin - dsimp, - rw category_theory.functor.map_id, - simp only [comp_id, add_neg_equiv_counit_iso_inv, eq_to_hom_refl, id_comp, nat_trans.comp_app, - discrete.functor_map_id, nat_trans.id_app, functor.map_neg, functor.map_comp, obj_μ_app, - obj_ε_inv_app, assoc, μ_naturality_assoc, neg_neg, μ_inv_hom_app_assoc], - erw [μ_inv_hom_app, category.comp_id, obj_zero_map_μ_app], - simp only [discrete.functor_map_id, nat_trans.id_app, comp_id, ε_hom_inv_app_assoc], - end } - -/-- -There is a natural transformation between the identity functor on triangles in `C`, -and the composition of an inverse rotation with a rotation. --/ -@[simps] -def inv_rot_comp_rot_inv : 𝟭 (triangle C) ⟶ inv_rotate C ⋙ rotate C := -{ app := to_rotate_inv_rotate, - naturality' := begin - introv, ext, - { dsimp, rw [comp_id, id_comp] }, - { dsimp, rw [comp_id, id_comp] }, - { dsimp, - rw [add_neg_equiv_counit_iso_inv, eq_to_hom_map, eq_to_hom_refl, id_comp], - simp only [nat_trans.comp_app, assoc], - erw [μ_inv_naturality, ε_naturality_assoc] }, - end } - -/-- -The natural transformations between the composition of a rotation with an inverse rotation -on triangles in `C`, and the identity functor on triangles are natural isomorphisms -(they are isomorphisms in the category of functors). --/ +/-- The counit isomorphism of the auto-equivalence of categories `triangle_rotation C` of +`triangle C` given by the rotation of triangles. -/ @[simps] def inv_rot_comp_rot : inv_rotate C ⋙ rotate C ≅ 𝟭 (triangle C) := -{ hom := inv_rot_comp_rot_hom, - inv := inv_rot_comp_rot_inv } +nat_iso.of_components (λ T, triangle.iso_mk _ _ (iso.refl _) (iso.refl _) + ((shift_equiv C (1 : ℤ)).counit_iso.app T.obj₃) (by tidy) (by tidy) (by tidy)) (by tidy) -variables (C) +variable (C) /-- Rotating triangles gives an auto-equivalence on the category of triangles in `C`. @@ -330,19 +140,7 @@ def triangle_rotation : equivalence (triangle C) (triangle C) := { functor := rotate C, inverse := inv_rotate C, unit_iso := rot_comp_inv_rot, - counit_iso := inv_rot_comp_rot, - functor_unit_iso_comp' := begin - introv, ext, - { dsimp, rw comp_id }, - { dsimp, rw comp_id }, - { dsimp, - rw unit_of_tensor_iso_unit_inv_app, - simp only [discrete.functor_map_id, nat_trans.id_app, id_comp, functor.map_comp, obj_ε_app, - obj_μ_inv_app, assoc, add_neg_equiv_counit_iso_hom, eq_to_hom_refl, nat_trans.comp_app, - ε_inv_app_obj, comp_id, μ_inv_hom_app_assoc], - erw [μ_inv_hom_app_assoc, μ_inv_hom_app], - refl } - end } + counit_iso := inv_rot_comp_rot, } variables {C} From 3e175454d8aa6a94734afcb2ae20a2f2b6660ea5 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 28 Mar 2023 23:45:20 +0000 Subject: [PATCH 0714/1291] feat(order/upper_lower/locally_finite): Upper closure preserves finiteness (#18678) `locally_finite_order` and `upper_set` don't naturally impot one another, so I'm dumping those two lemmas in a new file. --- src/order/upper_lower/locally_finite.lean | 26 +++++++++++++++++++++++ 1 file changed, 26 insertions(+) create mode 100644 src/order/upper_lower/locally_finite.lean diff --git a/src/order/upper_lower/locally_finite.lean b/src/order/upper_lower/locally_finite.lean new file mode 100644 index 0000000000000..4c8e171479c98 --- /dev/null +++ b/src/order/upper_lower/locally_finite.lean @@ -0,0 +1,26 @@ +/- +Copyright (c) 2023 Yaël Dillies, Sara Rousta. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies +-/ +import order.locally_finite +import order.upper_lower.basic + +/-! +# Upper and lower sets in a locally finite order + +In this file we characterise the interaction of `upper_set`/`lower_set` and `locally_finite_order`. +-/ + +namespace set +variables {α : Type*} [preorder α] {s : set α} + +protected lemma finite.upper_closure [locally_finite_order_top α] (hs : s.finite) : + (upper_closure s : set α).finite := +by { rw coe_upper_closure, exact hs.bUnion (λ _ _, finite_Ici _) } + +protected lemma finite.lower_closure [locally_finite_order_bot α] (hs : s.finite) : + (lower_closure s : set α).finite := +by { rw coe_lower_closure, exact hs.bUnion (λ _ _, finite_Iic _) } + +end set From 95127066efeffbd0ceddbda2e5cd8f755e74dd2e Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Wed, 29 Mar 2023 02:46:55 +0000 Subject: [PATCH 0715/1291] chore(scripts/list-attributes.sh): script to track attributes semireducible/irreducible (#18165) This script generates the data for #18164 --- scripts/list-attributes.sh | 15 +++++++++++++++ 1 file changed, 15 insertions(+) create mode 100755 scripts/list-attributes.sh diff --git a/scripts/list-attributes.sh b/scripts/list-attributes.sh new file mode 100755 index 0000000000000..eb83da538a10c --- /dev/null +++ b/scripts/list-attributes.sh @@ -0,0 +1,15 @@ +#!/usr/bin/env bash + +# This script generates the data for mathlib#18164 + +URL_BASE="https://github.com/leanprover-community/mathlib/blob" +SHA=$(git rev-parse HEAD) + +IFS=":" +git grep -n "local attribute \[semireducible\]\|attribute \[irreducible\]" | \ + grep -v 'src/tactic\|src/meta\|test' | \ + while read fn ln rest; do + grep --silent "SYNCHRONIZED WITH MATHLIB4" $fn || \ + echo "$URL_BASE/$SHA/$fn#L$ln" + done + From a7e36e48519ab281320c4d192da6a7b348ce40ad Mon Sep 17 00:00:00 2001 From: arienmalec Date: Wed, 29 Mar 2023 02:46:56 +0000 Subject: [PATCH 0716/1291] refactor(data/seq): scope seq and wseq to namespace stream (#18284) Adds namespace `stream` to `seq` and `wseq` to ease the Mathlib4 port (as `Seq` now name clashes with class `Seq` (`has_seq` in Lean 3). This requires disambiguation between `stream` and `computation` where lemmas in each namespace have the same name, and requires explicit scoping to mathlib files that reference `seq` and `wseq` (often by `open stream.seq as seq`) Co-authored-by: Scott Morrison --- docs/overview.yaml | 6 +- src/algebra/continued_fractions/basic.lean | 1 + .../computation/approximations.lean | 4 +- .../computation/basic.lean | 8 +- .../computation/terminates_iff_rat.lean | 1 + .../computation/translations.lean | 3 +- .../convergents_equiv.lean | 1 + .../terminated_stable.lean | 1 + .../continued_fractions/translations.lean | 1 + src/data/seq/parallel.lean | 11 +-- src/data/seq/seq.lean | 24 +++--- src/data/seq/wseq.lean | 74 +++++++++++-------- 12 files changed, 80 insertions(+), 55 deletions(-) diff --git a/docs/overview.yaml b/docs/overview.yaml index 1857a175084bf..f5178f64aff7e 100644 --- a/docs/overview.yaml +++ b/docs/overview.yaml @@ -138,7 +138,7 @@ General algebra: perfection of a ring: 'ring.perfection' Transcendental numbers: Liouville's theorem on existence of transcendental numbers: liouville.is_transcendental - + Representation theory: representation : 'representation' category of finite dimensional representations : 'fdRep' @@ -445,8 +445,8 @@ Data structures: difference list: 'dlist' lazy list: 'lazy_list' stream: 'stream' - sequence: 'seq' - weak sequence: 'wseq' + sequence: 'stream.seq' + weak sequence: 'stream.wseq' Sets: set: 'set' diff --git a/src/algebra/continued_fractions/basic.lean b/src/algebra/continued_fractions/basic.lean index 643a4d8a3a99c..dfb55dd60930f 100644 --- a/src/algebra/continued_fractions/basic.lean +++ b/src/algebra/continued_fractions/basic.lean @@ -47,6 +47,7 @@ variable (α : Type*) protected structure generalized_continued_fraction.pair := (a : α) (b : α) open generalized_continued_fraction +open stream.seq as seq /-! Interlude: define some expected coercions and instances. -/ namespace generalized_continued_fraction.pair diff --git a/src/algebra/continued_fractions/computation/approximations.lean b/src/algebra/continued_fractions/computation/approximations.lean index 99a04f2f26d94..19ca7909b95c2 100644 --- a/src/algebra/continued_fractions/computation/approximations.lean +++ b/src/algebra/continued_fractions/computation/approximations.lean @@ -296,9 +296,9 @@ theorem of_denom_mono : (of v).denominators n ≤ (of v).denominators (n + 1) := begin let g := of v, cases (decidable.em $ g.partial_denominators.terminated_at n) with terminated not_terminated, - { have : g.partial_denominators.nth n = none, by rwa seq.terminated_at at terminated, + { have : g.partial_denominators.nth n = none, by rwa stream.seq.terminated_at at terminated, have : g.terminated_at n, from - terminated_at_iff_part_denom_none.elim_right (by rwa seq.terminated_at at terminated), + terminated_at_iff_part_denom_none.elim_right (by rwa stream.seq.terminated_at at terminated), have : g.denominators (n + 1) = g.denominators n, from denominators_stable_of_terminated n.le_succ this, rw this }, diff --git a/src/algebra/continued_fractions/computation/basic.lean b/src/algebra/continued_fractions/computation/basic.lean index 61ac955115559..eef0b6a89ce59 100644 --- a/src/algebra/continued_fractions/computation/basic.lean +++ b/src/algebra/continued_fractions/computation/basic.lean @@ -131,6 +131,7 @@ protected def stream (v : K) : stream $ option (int_fract_pair K) | (n + 1) := (stream n).bind $ λ ap_n, if ap_n.fr = 0 then none else some (int_fract_pair.of ap_n.fr⁻¹) + /-- Shows that `int_fract_pair.stream` has the sequence property, that is once we return `none` at position `n`, we also return `none` at `n + 1`. @@ -147,10 +148,11 @@ This is just an intermediate representation and users should not (need to) direc it. The setup of rewriting/simplification lemmas that make the definitions easy to use is done in `algebra.continued_fractions.computation.translations`. -/ -protected def seq1 (v : K) : seq1 $ int_fract_pair K := +protected def seq1 (v : K) : stream.seq1 $ int_fract_pair K := ⟨ int_fract_pair.of v,--the head - seq.tail -- take the tail of `int_fract_pair.stream` since the first element is already in the - -- head create a sequence from `int_fract_pair.stream` + stream.seq.tail -- take the tail of `int_fract_pair.stream` since the first element is already in + -- the head + -- create a sequence from `int_fract_pair.stream` ⟨ int_fract_pair.stream v, -- the underlying stream @stream_is_seq _ _ _ v ⟩ ⟩ -- the proof that the stream is a sequence diff --git a/src/algebra/continued_fractions/computation/terminates_iff_rat.lean b/src/algebra/continued_fractions/computation/terminates_iff_rat.lean index 53e9232e5047a..d4977bd208740 100644 --- a/src/algebra/continued_fractions/computation/terminates_iff_rat.lean +++ b/src/algebra/continued_fractions/computation/terminates_iff_rat.lean @@ -28,6 +28,7 @@ rational, continued fraction, termination -/ namespace generalized_continued_fraction +open stream.seq as seq open generalized_continued_fraction (of) diff --git a/src/algebra/continued_fractions/computation/translations.lean b/src/algebra/continued_fractions/computation/translations.lean index 3d999f84a89ae..2d7683d10376a 100644 --- a/src/algebra/continued_fractions/computation/translations.lean +++ b/src/algebra/continued_fractions/computation/translations.lean @@ -38,6 +38,7 @@ of three sections: namespace generalized_continued_fraction open generalized_continued_fraction (of) +open stream.seq as seq /- Fix a discrete linear ordered floor field and a value `v`. -/ variables {K : Type*} [linear_ordered_field K] [floor_ring K] {v : K} @@ -193,7 +194,7 @@ option.map_eq_none lemma of_terminated_at_n_iff_succ_nth_int_fract_pair_stream_eq_none : (of v).terminated_at n ↔ int_fract_pair.stream v (n + 1) = none := -by rw [of_terminated_at_iff_int_fract_pair_seq1_terminated_at, seq.terminated_at, +by rw [of_terminated_at_iff_int_fract_pair_seq1_terminated_at, stream.seq.terminated_at, int_fract_pair.nth_seq1_eq_succ_nth_stream] end termination diff --git a/src/algebra/continued_fractions/convergents_equiv.lean b/src/algebra/continued_fractions/convergents_equiv.lean index 244c9e843532a..bc5d40a4f6e8b 100644 --- a/src/algebra/continued_fractions/convergents_equiv.lean +++ b/src/algebra/continued_fractions/convergents_equiv.lean @@ -65,6 +65,7 @@ fractions, recurrence, equivalence variables {K : Type*} {n : ℕ} namespace generalized_continued_fraction +open stream.seq as seq variables {g : generalized_continued_fraction K} {s : seq $ pair K} diff --git a/src/algebra/continued_fractions/terminated_stable.lean b/src/algebra/continued_fractions/terminated_stable.lean index 0bdcae77dd778..097b544df7a02 100644 --- a/src/algebra/continued_fractions/terminated_stable.lean +++ b/src/algebra/continued_fractions/terminated_stable.lean @@ -13,6 +13,7 @@ We show that the continuants and convergents of a gcf stabilise once the gcf ter -/ namespace generalized_continued_fraction +open stream.seq as seq variables {K : Type*} {g : generalized_continued_fraction K} {n m : ℕ} diff --git a/src/algebra/continued_fractions/translations.lean b/src/algebra/continued_fractions/translations.lean index 3a1f74dc5a8dc..dc8bfb1a2c62b 100644 --- a/src/algebra/continued_fractions/translations.lean +++ b/src/algebra/continued_fractions/translations.lean @@ -14,6 +14,7 @@ Some simple translation lemmas between the different definitions of functions de -/ namespace generalized_continued_fraction +open stream.seq as seq section general /-! diff --git a/src/data/seq/parallel.lean b/src/data/seq/parallel.lean index 80f7dfc0832e7..38b38a046354b 100644 --- a/src/data/seq/parallel.lean +++ b/src/data/seq/parallel.lean @@ -14,7 +14,8 @@ import data.seq.wseq universes u v namespace computation -open wseq +open stream.wseq as wseq +open stream.seq as seq variables {α : Type u} {β : Type v} def parallel.aux2 : list (computation α) → α ⊕ list (computation α) := @@ -26,7 +27,7 @@ end) (sum.inr []) def parallel.aux1 : list (computation α) × wseq (computation α) → α ⊕ list (computation α) × wseq (computation α) | (l, S) := rmap (λ l', match seq.destruct S with - | none := (l', nil) + | none := (l', seq.nil) | some (none, S') := (l', S') | some (some c, S') := (c::l', S') end) (parallel.aux2 l) @@ -156,7 +157,7 @@ begin exact ⟨c, or.inl cl, ac⟩ }, { induction e : seq.destruct S with a; rw e at h', { exact let ⟨d, o, ad⟩ := IH _ _ h', - ⟨c, cl, ac⟩ := this a ⟨d, o.resolve_right (not_mem_nil _), ad⟩ in + ⟨c, cl, ac⟩ := this a ⟨d, o.resolve_right (wseq.not_mem_nil _), ad⟩ in ⟨c, or.inl cl, ac⟩ }, { cases a with o S', cases o with c; simp [parallel.aux1] at h'; rcases IH _ _ h' with ⟨d, dl | dS', ad⟩, @@ -196,8 +197,8 @@ theorem parallel_empty (S : wseq (computation α)) (h : S.head ~> none) : parallel S = empty _ := eq_empty_of_not_terminates $ λ ⟨⟨a, m⟩⟩, let ⟨c, cs, ac⟩ := exists_of_mem_parallel m, - ⟨n, nm⟩ := exists_nth_of_mem cs, - ⟨c', h'⟩ := head_some_of_nth_some nm in by injection h h' + ⟨n, nm⟩ := wseq.exists_nth_of_mem cs, + ⟨c', h'⟩ := wseq.head_some_of_nth_some nm in by injection h h' -- The reason this isn't trivial from exists_of_mem_parallel is because it eliminates to Sort def parallel_rec {S : wseq (computation α)} (C : α → Sort v) diff --git a/src/data/seq/seq.lean b/src/data/seq/seq.lean index 4571c0c1a4020..182edc11ec884 100644 --- a/src/data/seq/seq.lean +++ b/src/data/seq/seq.lean @@ -9,6 +9,7 @@ import data.nat.basic import data.stream.init import data.seq.computation +namespace stream universes u v w /- @@ -20,7 +21,7 @@ coinductive seq (α : Type u) : Type u /-- A stream `s : option α` is a sequence if `s.nth n = none` implies `s.nth (n + 1) = none`. -/ -def stream.is_seq {α : Type u} (s : stream (option α)) : Prop := +def is_seq {α : Type u} (s : stream (option α)) : Prop := ∀ {n : ℕ}, s n = none → s (n + 1) = none /-- `seq α` is the type of possibly infinite lists (referred here as sequences). @@ -95,6 +96,7 @@ def head (s : seq α) : option α := nth s 0 /-- Get the tail of a sequence (or `nil` if the sequence is `nil`) -/ def tail (s : seq α) : seq α := ⟨s.1.tail, λ n, by { cases s with f al, exact al }⟩ +/-- member definition for `seq`-/ protected def mem (a : α) (s : seq α) := some a ∈ s.1 instance : has_mem α (seq α) := @@ -209,6 +211,7 @@ begin apply h1 _ _ (or.inr (IH e)) } end +/-- Corecursor over pairs of `option` values-/ def corec.F (f : β → option (α × β)) : option β → option α × option β | none := (none, none) | (some b) := match f b with none := (none, none) | some (a, b') := (some a, some b') end @@ -252,12 +255,14 @@ section bisim local infix (name := R) ` ~ `:50 := R + /-- Bisimilarity relation over `option` of `seq1 α`-/ def bisim_o : option (seq1 α) → option (seq1 α) → Prop | none none := true | (some (a, s)) (some (a', s')) := a = a' ∧ R s s' | _ _ := false attribute [simp] bisim_o + /-- a relation is bisimiar if it meets the `bisim_o` test-/ def is_bisimulation := ∀ ⦃s₁ s₂⦄, s₁ ~ s₂ → bisim_o R (destruct s₁) (destruct s₂) -- If two streams are bisimilar, then they are equal @@ -667,11 +672,11 @@ end seq namespace seq1 variables {α : Type u} {β : Type v} {γ : Type w} -open seq +open stream.seq /-- Convert a `seq1` to a sequence. -/ def to_seq : seq1 α → seq α -| (a, s) := cons a s +| (a, s) := seq.cons a s instance coe_seq : has_coe (seq1 α) (seq α) := ⟨to_seq⟩ @@ -685,13 +690,13 @@ theorem map_id : ∀ (s : seq1 α), map id s = s | ⟨a, s⟩ := by simp [map] def join : seq1 (seq1 α) → seq1 α | ((a, s), S) := match destruct s with | none := (a, seq.join S) - | some s' := (a, seq.join (cons s' S)) + | some s' := (a, seq.join (seq.cons s' S)) end @[simp] theorem join_nil (a : α) (S) : join ((a, nil), S) = (a, seq.join S) := rfl @[simp] theorem join_cons (a b : α) (s S) : - join ((a, cons b s), S) = (a, seq.join (cons (b, s) S)) := + join ((a, seq.cons b s), S) = (a, seq.join (seq.cons (b, s) S)) := by dsimp [join]; rw [destruct_cons]; refl /-- The `return` operator for the `seq1` monad, @@ -726,8 +731,8 @@ end @[simp] theorem map_join' (f : α → β) (S) : seq.map f (seq.join S) = seq.join (seq.map (map f) S) := begin - apply eq_of_bisim (λ s1 s2, - ∃ s S, s1 = append s (seq.map f (seq.join S)) ∧ + apply seq.eq_of_bisim (λ s1 s2, + ∃ s S, s1 = seq.append s (seq.map f (seq.join S)) ∧ s2 = append s (seq.join (seq.map (map f) S))), { intros s1 s2 h, exact match s1, s2, h with ._, ._, ⟨s, S, rfl, rfl⟩ := begin apply rec_on s; simp, @@ -745,7 +750,7 @@ end @[simp] theorem join_join (SS : seq (seq1 (seq1 α))) : seq.join (seq.join SS) = seq.join (seq.map join SS) := begin - apply eq_of_bisim (λ s1 s2, + apply seq.eq_of_bisim (λ s1 s2, ∃ s SS, s1 = seq.append s (seq.join (seq.join SS)) ∧ s2 = seq.append s (seq.join (seq.map join SS))), { intros s1 s2 h, exact match s1, s2, h with ._, ._, ⟨s, SS, rfl, rfl⟩ := begin @@ -755,7 +760,7 @@ begin apply rec_on s; simp, { exact ⟨_, _, rfl, rfl⟩ }, { intros x s, - refine ⟨cons x (append s (seq.join S)), SS, _, _⟩; simp } } }, + refine ⟨seq.cons x (append s (seq.join S)), SS, _, _⟩; simp } } }, { intros x s, exact ⟨s, SS, rfl, rfl⟩ } end end }, { refine ⟨nil, SS, _, _⟩; simp } @@ -788,3 +793,4 @@ instance : is_lawful_monad seq1 := bind_assoc := @bind_assoc } end seq1 +end stream diff --git a/src/data/seq/wseq.lean b/src/data/seq/wseq.lean index 4182c7fe937a7..c55be3a1b2301 100644 --- a/src/data/seq/wseq.lean +++ b/src/data/seq/wseq.lean @@ -6,6 +6,7 @@ Authors: Mario Carneiro import data.list.basic import data.seq.seq +namespace stream open function universes u v w @@ -71,6 +72,7 @@ def rec_on {C : wseq α → Sort v} (s : wseq α) (h1 : C nil) (h2 : ∀ x s, C (cons x s)) (h3 : ∀ s, C (think s)) : C s := seq.rec_on s h1 (λ o, option.rec_on o h3 h2) +/-- membership for weak sequences-/ protected def mem (a : α) (s : wseq α) := seq.mem (some a) s instance : has_mem α (wseq α) := @@ -324,6 +326,7 @@ seq.join ((λ o : option (wseq α), match o with def bind (s : wseq α) (f : α → wseq β) : wseq β := join (map f s) +/-- lift a relation to a relation over weak sequences -/ @[simp] def lift_rel_o (R : α → β → Prop) (C : wseq α → wseq β → Prop) : option (α × wseq α) → option (β × wseq β) → Prop | none none := true @@ -342,6 +345,7 @@ theorem lift_rel_o.imp_right (R : α → β → Prop) {C D : wseq α → wseq β (H : ∀ s t, C s t → D s t) {o p} : lift_rel_o R C o p → lift_rel_o R D o p := lift_rel_o.imp (λ _ _, id) H +/-- Definitino of bisimilarity for weak sequences-/ @[simp] def bisim_o (R : wseq α → wseq α → Prop) : option (α × wseq α) → option (α × wseq α) → Prop := lift_rel_o (=) R @@ -557,6 +561,7 @@ by { simp [think, join], unfold functor.map, simp [join, cons, append] } @[simp] theorem append_assoc (s t u : wseq α) : append (append s t) u = append s (append t u) := seq.append_assoc _ _ _ +/-- auxilary defintion of tail over weak sequences-/ @[simp] def tail.aux : option (α × wseq α) → computation (option (α × wseq α)) | none := return none | (some (a, s)) := destruct s @@ -569,6 +574,7 @@ begin apply (@pure_bind computation _ _ _ _ _ _).trans _; simp end +/-- auxilary defintion of drop over weak sequences-/ @[simp] def drop.aux : ℕ → option (α × wseq α) → computation (option (α × wseq α)) | 0 := return | (n+1) := λ a, tail.aux a >>= drop.aux n @@ -591,7 +597,7 @@ theorem head_terminates_of_head_tail_terminates (s : wseq α) [T : terminates (h simp [tail] at h, rcases exists_of_mem_bind h with ⟨s', h1, h2⟩, unfold functor.map at h1, - exact let ⟨t, h3, h4⟩ := exists_of_mem_map h1 in terminates_of_mem h3 + exact let ⟨t, h3, h4⟩ := computation.exists_of_mem_map h1 in computation.terminates_of_mem h3 end theorem destruct_some_of_destruct_tail_some {s : wseq α} {a} @@ -599,7 +605,7 @@ theorem destruct_some_of_destruct_tail_some {s : wseq α} {a} begin unfold tail functor.map at h, simp at h, rcases exists_of_mem_bind h with ⟨t, tm, td⟩, clear h, - rcases exists_of_mem_map tm with ⟨t', ht', ht2⟩, clear tm, + rcases computation.exists_of_mem_map tm with ⟨t', ht', ht2⟩, clear tm, cases t' with t'; rw ←ht2 at td; simp at td, { have := mem_unique td (ret_mem _), contradiction }, { exact ⟨_, ht'⟩ } @@ -609,10 +615,10 @@ theorem head_some_of_head_tail_some {s : wseq α} {a} (h : some a ∈ head (tail s)) : ∃ a', some a' ∈ head s := begin unfold head at h, - rcases exists_of_mem_map h with ⟨o, md, e⟩, clear h, + rcases computation.exists_of_mem_map h with ⟨o, md, e⟩, clear h, cases o with o; injection e with h', clear e h', cases destruct_some_of_destruct_tail_some md with a am, - exact ⟨_, mem_map ((<$>) (@prod.fst α (wseq α))) am⟩ + exact ⟨_, computation.mem_map ((<$>) (@prod.fst α (wseq α))) am⟩ end theorem head_some_of_nth_some {s : wseq α} {a n} @@ -719,7 +725,7 @@ theorem mem_of_mem_dropn {s : wseq α} {a} : ∀ {n}, a ∈ drop s n → a ∈ s theorem nth_mem {s : wseq α} {a n} : some a ∈ nth s n → a ∈ s := begin revert s, induction n with n IH; intros s h, - { rcases exists_of_mem_map h with ⟨o, h1, h2⟩, + { rcases computation.exists_of_mem_map h with ⟨o, h1, h2⟩, cases o with o; injection h2 with h', cases o with a' s', exact (eq_or_mem_iff_mem h1).2 (or.inl h'.symm) }, @@ -742,7 +748,7 @@ theorem exists_dropn_of_mem {s : wseq α} {a} (h : a ∈ s) : ∃ n s', some (a, s') ∈ destruct (drop s n) := let ⟨n, h⟩ := exists_nth_of_mem h in ⟨n, begin rcases (head_terminates_iff _).1 ⟨⟨_, h⟩⟩ with ⟨⟨o, om⟩⟩, - have := mem_unique (mem_map _ om) h, + have := computation.mem_unique (computation.mem_map _ om) h, cases o with o; injection this with i, cases o with a' s', dsimp at i, rw i at om, exact ⟨_, om⟩ @@ -765,9 +771,9 @@ end theorem exists_of_lift_rel_left {R : α → β → Prop} {s t} (H : lift_rel R s t) {a} (h : a ∈ s) : ∃ {b}, b ∈ t ∧ R a b := let ⟨n, h⟩ := exists_nth_of_mem h, - ⟨some (._, s'), sd, rfl⟩ := exists_of_mem_map h, + ⟨some (._, s'), sd, rfl⟩ := computation.exists_of_mem_map h, ⟨some (b, t'), td, ⟨ab, _⟩⟩ := (lift_rel_dropn_destruct H n).left sd in -⟨b, nth_mem (mem_map ((<$>) prod.fst.{v v}) td), ab⟩ +⟨b, nth_mem (computation.mem_map ((<$>) prod.fst.{v v}) td), ab⟩ theorem exists_of_lift_rel_right {R : α → β → Prop} {s t} (H : lift_rel R s t) {b} (h : b ∈ t) : ∃ {a}, a ∈ s ∧ R a b := @@ -807,7 +813,7 @@ by unfold equiv; simp; exact h theorem think_equiv (s : wseq α) : think s ~ s := by unfold equiv; simp; apply equiv.refl -theorem think_congr {s t : wseq α} (a : α) (h : s ~ t) : think s ~ think t := +theorem think_congr {s t : wseq α} (h : s ~ t) : think s ~ think t := by unfold equiv; simp; exact h theorem head_congr : ∀ {s t : wseq α}, s ~ t → head s ~ head t := @@ -820,11 +826,11 @@ begin cases destruct_congr h with l r, rcases l dsm with ⟨dt, dtm, dst⟩, cases ds with a; cases dt with b, - { apply mem_map _ dtm }, + { apply computation.mem_map _ dtm }, { cases b, cases dst }, { cases a, cases dst }, { cases a with a s', cases b with b t', rw dst.left, - exact @mem_map _ _ (@functor.map _ _ (α × wseq α) _ prod.fst) + exact @computation.mem_map _ _ (@functor.map _ _ (α × wseq α) _ prod.fst) _ (destruct t) dtm } end @@ -885,23 +891,24 @@ theorem equiv.ext {s t : wseq α} (h : ∀ n, nth s n ~ nth t n) : s ~ t := { intros a b ma mb, cases a with a; cases b with b, { trivial }, - { injection mem_unique (mem_map _ ma) ((h 0 _).2 (mem_map _ mb)) }, - { injection mem_unique (mem_map _ ma) ((h 0 _).2 (mem_map _ mb)) }, + { injection mem_unique (computation.mem_map _ ma) ((h 0 _).2 (computation.mem_map _ mb)) }, + { injection mem_unique (computation.mem_map _ ma) ((h 0 _).2 (computation.mem_map _ mb)) }, { cases a with a s', cases b with b t', - injection mem_unique (mem_map _ ma) ((h 0 _).2 (mem_map _ mb)) with ab, + injection mem_unique + (computation.mem_map _ ma) ((h 0 _).2 (computation.mem_map _ mb)) with ab, refine ⟨ab, λ n, _⟩, - refine (nth_congr (flatten_equiv (mem_map _ ma)) n).symm.trans + refine (nth_congr (flatten_equiv (computation.mem_map _ ma)) n).symm.trans ((_ : nth (tail s) n ~ nth (tail t) n).trans - (nth_congr (flatten_equiv (mem_map _ mb)) n)), + (nth_congr (flatten_equiv (computation.mem_map _ mb)) n)), rw [nth_tail, nth_tail], apply h } } end⟩ theorem length_eq_map (s : wseq α) : length s = computation.map list.length (to_list s) := begin - refine eq_of_bisim + refine computation.eq_of_bisim (λ c1 c2, ∃ (l : list α) (s : wseq α), - c1 = corec length._match_2 (l.length, s) ∧ - c2 = computation.map list.length (corec to_list._match_2 (l, s))) + c1 = computation.corec length._match_2 (l.length, s) ∧ + c2 = computation.map list.length (computation.corec to_list._match_2 (l, s))) _ ⟨[], s, rfl, rfl⟩, intros s1 s2 h, rcases h with ⟨l, s, h⟩, rw [h.left, h.right], apply s.rec_on _ (λ a s, _) (λ s, _); @@ -918,27 +925,27 @@ show seq.map some (seq.of_list (a :: l)) = seq.cons (some a) (seq.map some (seq.of_list l)), by simp @[simp] theorem to_list'_nil (l : list α) : - corec to_list._match_2 (l, nil) = return l.reverse := + computation.corec to_list._match_2 (l, nil) = return l.reverse := destruct_eq_ret rfl @[simp] theorem to_list'_cons (l : list α) (s : wseq α) (a : α) : - corec to_list._match_2 (l, cons a s) = - (corec to_list._match_2 (a::l, s)).think := + computation.corec to_list._match_2 (l, cons a s) = + (computation.corec to_list._match_2 (a::l, s)).think := destruct_eq_think $ by simp [to_list, cons] @[simp] theorem to_list'_think (l : list α) (s : wseq α) : - corec to_list._match_2 (l, think s) = - (corec to_list._match_2 (l, s)).think := + computation.corec to_list._match_2 (l, think s) = + (computation.corec to_list._match_2 (l, s)).think := destruct_eq_think $ by simp [to_list, think] theorem to_list'_map (l : list α) (s : wseq α) : - corec to_list._match_2 (l, s) = + computation.corec to_list._match_2 (l, s) = ((++) l.reverse) <$> to_list s := begin - refine eq_of_bisim + refine computation.eq_of_bisim (λ c1 c2, ∃ (l' : list α) (s : wseq α), - c1 = corec to_list._match_2 (l' ++ l, s) ∧ - c2 = computation.map ((++) l.reverse) (corec to_list._match_2 (l', s))) + c1 = computation.corec to_list._match_2 (l' ++ l, s) ∧ + c2 = computation.map ((++) l.reverse) (computation.corec to_list._match_2 (l', s))) _ ⟨[], s, rfl, rfl⟩, intros s1 s2 h, rcases h with ⟨l', s, h⟩, rw [h.left, h.right], apply s.rec_on _ (λ a s, _) (λ s, _); @@ -955,7 +962,7 @@ destruct_eq_think $ by unfold to_list; simp; rw to_list'_map; simp; refl destruct_eq_ret rfl theorem to_list_of_list (l : list α) : l ∈ to_list (of_list l) := -by induction l with a l IH; simp [ret_mem]; exact think_mem (mem_map _ IH) +by induction l with a l IH; simp [ret_mem]; exact think_mem (computation.mem_map _ IH) @[simp] theorem destruct_of_seq (s : seq α) : destruct (of_seq s) = return (s.head.map $ λ a, (a, of_seq s.tail)) := @@ -1061,7 +1068,7 @@ let ⟨t, tm, bt⟩ := exists_of_mem_join h, theorem destruct_map (f : α → β) (s : wseq α) : destruct (map f s) = computation.map (option.map (prod.map f (map f))) (destruct s) := begin - apply eq_of_bisim (λ c1 c2, ∃ s, c1 = destruct (map f s) ∧ + apply computation.eq_of_bisim (λ c1 c2, ∃ s, c1 = destruct (map f s) ∧ c2 = computation.map (option.map (prod.map f (map f))) (destruct s)), { intros c1 c2 h, cases h with s h, rw [h.left, h.right], apply s.rec_on _ (λ a s, _) (λ s, _); simp, @@ -1089,6 +1096,7 @@ end end⟩ theorem map_congr (f : α → β) {s t : wseq α} (h : s ~ t) : map f s ~ map f t := lift_rel_map _ _ h (λ _ _, congr_arg _) +/-- auxilary defintion of `destruct_append` over weak sequences-/ @[simp] def destruct_append.aux (t : wseq α) : option (α × wseq α) → computation (option (α × wseq α)) | none := destruct t @@ -1097,7 +1105,7 @@ lift_rel_map _ _ h (λ _ _, congr_arg _) theorem destruct_append (s t : wseq α) : destruct (append s t) = (destruct s).bind (destruct_append.aux t) := begin - apply eq_of_bisim (λ c1 c2, ∃ s t, c1 = destruct (append s t) ∧ + apply computation.eq_of_bisim (λ c1 c2, ∃ s t, c1 = destruct (append s t) ∧ c2 = (destruct s).bind (destruct_append.aux t)) _ ⟨s, t, rfl, rfl⟩, intros c1 c2 h, rcases h with ⟨s, t, h⟩, rw [h.left, h.right], apply s.rec_on _ (λ a s, _) (λ s, _); simp, @@ -1106,6 +1114,7 @@ begin { exact ⟨s, t, rfl, rfl⟩ } end +/-- auxilary defintion of `destruct_join` over weak sequences-/ @[simp] def destruct_join.aux : option (wseq α × wseq (wseq α)) → computation (option (α × wseq α)) | none := return none | (some (s, S)) := (destruct (append s (join S))).think @@ -1113,7 +1122,7 @@ end theorem destruct_join (S : wseq (wseq α)) : destruct (join S) = (destruct S).bind destruct_join.aux := begin - apply eq_of_bisim (λ c1 c2, c1 = c2 ∨ ∃ S, c1 = destruct (join S) ∧ + apply computation.eq_of_bisim (λ c1 c2, c1 = c2 ∨ ∃ S, c1 = destruct (join S) ∧ c2 = (destruct S).bind destruct_join.aux) _ (or.inr ⟨S, rfl, rfl⟩), intros c1 c2 h, exact match c1, c2, h with | _, _, (or.inl $ eq.refl c) := by cases c.destruct; simp @@ -1359,3 +1368,4 @@ instance : is_lawful_monad wseq := -/ end wseq +end stream From 98bbc3526516bca903bff09ea10c4206bf079e6b Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Wed, 29 Mar 2023 02:46:57 +0000 Subject: [PATCH 0717/1291] feat(set_theory/zfc/ordinal): more lemmas on transitive sets (#18307) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Yaël Dillies --- src/set_theory/zfc/basic.lean | 13 +++++++++++++ src/set_theory/zfc/ordinal.lean | 13 +++++++++++++ 2 files changed, 26 insertions(+) diff --git a/src/set_theory/zfc/basic.lean b/src/set_theory/zfc/basic.lean index ff429736a9045..8a27059edaabc 100644 --- a/src/set_theory/zfc/basic.lean +++ b/src/set_theory/zfc/basic.lean @@ -720,6 +720,19 @@ by { rw ←mem_to_set, simp } @[simp] theorem mem_diff {x y z : Set.{u}} : z ∈ x \ y ↔ z ∈ x ∧ z ∉ y := @@mem_sep (λ z : Set.{u}, z ∉ y) +@[simp] theorem sUnion_pair {x y : Set.{u}} : ⋃₀ ({x, y} : Set.{u}) = x ∪ y := +begin + ext, + simp_rw [mem_union, mem_sUnion, mem_pair], + split, + { rintro ⟨w, (rfl | rfl), hw⟩, + { exact or.inl hw }, + { exact or.inr hw } }, + { rintro (hz | hz), + { exact ⟨x, or.inl rfl, hz⟩ }, + { exact ⟨y, or.inr rfl, hz⟩ } } +end + theorem mem_wf : @well_founded Set (∈) := well_founded_lift₂_iff.mpr pSet.mem_wf diff --git a/src/set_theory/zfc/ordinal.lean b/src/set_theory/zfc/ordinal.lean index c752b5c314a48..0ff3f4855fc73 100644 --- a/src/set_theory/zfc/ordinal.lean +++ b/src/set_theory/zfc/ordinal.lean @@ -60,6 +60,19 @@ theorem is_transitive.sUnion' (H : ∀ y ∈ x, is_transitive y) : (⋃₀ x).is exact mem_sUnion_of_mem ((H w hw).mem_trans hz hw') hw end +protected theorem is_transitive.union (hx : x.is_transitive) (hy : y.is_transitive) : + (x ∪ y).is_transitive := +begin + rw ←sUnion_pair, + apply is_transitive.sUnion' (λ z, _), + rw mem_pair, + rintro (rfl | rfl), + assumption' +end + +protected theorem is_transitive.powerset (h : x.is_transitive) : (powerset x).is_transitive := +λ y hy z hz, by { rw mem_powerset at ⊢ hy, exact h.subset_of_mem (hy hz) } + theorem is_transitive_iff_sUnion_subset : x.is_transitive ↔ ⋃₀ x ⊆ x := ⟨λ h y hy, by { rcases mem_sUnion.1 hy with ⟨z, hz, hz'⟩, exact h.mem_trans hz' hz }, λ H y hy z hz, H $ mem_sUnion_of_mem hz hy⟩ From d3acee0d776b15ffb8318f327325ff343cc8bdcc Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Wed, 29 Mar 2023 02:46:58 +0000 Subject: [PATCH 0718/1291] feat(ring_theory/ideal/norm): add ideal.finite_set_of_abs_norm_eq (#18569) Prove that there are only finitely many ideals of `abs_norm` equal to $n$ for $n$ positive. --- src/ring_theory/ideal/norm.lean | 26 +++++++++++++++++++ .../ideal/quotient_operations.lean | 4 +++ 2 files changed, 30 insertions(+) diff --git a/src/ring_theory/ideal/norm.lean b/src/ring_theory/ideal/norm.lean index 158190d30f825..5311fe018ee5b 100644 --- a/src/ring_theory/ideal/norm.lean +++ b/src/ring_theory/ideal/norm.lean @@ -274,6 +274,10 @@ by rw [← ideal.one_eq_top, _root_.map_one] @[simp] lemma abs_norm_eq_one_iff {I : ideal S} : abs_norm I = 1 ↔ I = ⊤ := by rw [abs_norm_apply, card_quot_eq_one_iff] +lemma abs_norm_ne_zero_iff (I : ideal S) : ideal.abs_norm I ≠ 0 ↔ finite (S ⧸ I) := +⟨λ h,nat.finite_of_card_ne_zero h, + λ h, (@add_subgroup.finite_index_of_finite_quotient _ _ _ h).finite_index⟩ + /-- Let `e : S ≃ I` be an additive isomorphism (therefore a `ℤ`-linear equiv). Then an alternative way to compute the norm of `I` is given by taking the determinant of `e`. See `nat_abs_det_basis_change` for a more familiar formulation of this result. -/ @@ -424,6 +428,28 @@ lemma abs_norm_mem (I : ideal S) : ↑I.abs_norm ∈ I := by rw [abs_norm_apply, card_quot, ← ideal.quotient.eq_zero_iff_mem, map_nat_cast, quotient.index_eq_zero] +lemma span_singleton_abs_norm_le (I : ideal S) : + ideal.span { (ideal.abs_norm I : S) } ≤ I := +by simp only [ideal.span_le, set.singleton_subset_iff, set_like.mem_coe, ideal.abs_norm_mem I] + +lemma finite_set_of_abs_norm_eq [char_zero S] {n : ℕ} (hn : 0 < n) : + { I : ideal S | ideal.abs_norm I = n }.finite := +begin + let f := λ I : ideal S, ideal.map (ideal.quotient.mk (@ideal.span S _ {n})) I, + refine @set.finite.of_finite_image _ _ _ f _ _, + { suffices : finite (S ⧸ @ideal.span S _ {n}), + { let g := (coe : ideal (S ⧸ @ideal.span S _ {n}) → set (S ⧸ @ideal.span S _ {n})), + refine @set.finite.of_finite_image _ _ _ g _ (set_like.coe_injective.inj_on _), + exact set.finite.subset (@set.finite_univ _ (@set.finite' _ this)) ( set.subset_univ _), }, + rw [← abs_norm_ne_zero_iff, abs_norm_span_singleton], + simpa only [ne.def, int.nat_abs_eq_zero, algebra.norm_eq_zero_iff, nat.cast_eq_zero] + using ne_of_gt hn, }, + { intros I hI J hJ h, + rw [← comap_map_mk (span_singleton_abs_norm_le I), ← hI.symm, + ← comap_map_mk (span_singleton_abs_norm_le J), ← hJ.symm], + exact congr_arg (ideal.comap (ideal.quotient.mk (@ideal.span S _ {n}))) h, }, +end + end ideal end ring_of_integers diff --git a/src/ring_theory/ideal/quotient_operations.lean b/src/ring_theory/ideal/quotient_operations.lean index 0142d188c484e..4bbea539b3a0e 100644 --- a/src/ring_theory/ideal/quotient_operations.lean +++ b/src/ring_theory/ideal/quotient_operations.lean @@ -126,6 +126,10 @@ lemma mem_quotient_iff_mem {I J : ideal R} (hIJ : I ≤ J) {x : R} : quotient.mk I x ∈ J.map (quotient.mk I) ↔ x ∈ J := by rw [mem_quotient_iff_mem_sup, sup_eq_left.mpr hIJ] +lemma comap_map_mk {I J : ideal R} (h : I ≤ J) : + ideal.comap (ideal.quotient.mk I) (ideal.map (ideal.quotient.mk I) J) = J := +by { ext, rw [← ideal.mem_quotient_iff_mem h, ideal.mem_comap], } + section quotient_algebra variables (R₁ R₂ : Type*) {A B : Type*} From 200eda15d8ff5669854ff6bcc10aaf37cb70498f Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Wed, 29 Mar 2023 04:43:47 +0000 Subject: [PATCH 0719/1291] feat(category_theory/idempotents): idempotent completeness of homological_complex (#17921) --- src/algebra/homology/additive.lean | 37 +++++++++++++++++++ .../idempotents/homological_complex.lean | 13 ++++++- src/category_theory/idempotents/karoubi.lean | 8 ++++ 3 files changed, 57 insertions(+), 1 deletion(-) diff --git a/src/algebra/homology/additive.lean b/src/algebra/homology/additive.lean index 2d2d5662d2939..9ecd5b5fab41d 100644 --- a/src/algebra/homology/additive.lean +++ b/src/algebra/homology/additive.lean @@ -107,6 +107,17 @@ def functor.map_homological_complex (F : V ⥤ W) [F.additive] (c : complex_shap { f := λ i, F.map (f.f i), comm' := λ i j h, by { dsimp, rw [←F.map_comp, ←F.map_comp, f.comm], }, }, }. +variable (V) + +/-- The functor on homological complexes induced by the identity functor is +isomorphic to the identity functor. -/ +@[simps] +def functor.map_homological_complex_id_iso (c : complex_shape ι) : + (𝟭 V).map_homological_complex c ≅ 𝟭 _ := +nat_iso.of_components (λ K, hom.iso_of_components (λ i, iso.refl _) (by tidy)) (by tidy) + +variable {V} + instance functor.map_homogical_complex_additive (F : V ⥤ W) [F.additive] (c : complex_shape ι) : (F.map_homological_complex c).additive := {} @@ -147,6 +158,32 @@ by tidy (nat_trans.map_homological_complex α c).app C ≫ (G.map_homological_complex c).map f := by tidy +/-- +A natural isomorphism between functors induces a natural isomorphism +between those functors applied to homological complexes. +-/ +@[simps] +def nat_iso.map_homological_complex {F G : V ⥤ W} [F.additive] [G.additive] + (α : F ≅ G) (c : complex_shape ι) : F.map_homological_complex c ≅ G.map_homological_complex c := +{ hom := α.hom.map_homological_complex c, + inv := α.inv.map_homological_complex c, + hom_inv_id' := by simpa only [← nat_trans.map_homological_complex_comp, α.hom_inv_id], + inv_hom_id' := by simpa only [← nat_trans.map_homological_complex_comp, α.inv_hom_id], } + +/-- +An equivalence of categories induces an equivalences between the respective categories +of homological complex. +-/ +@[simps] +def equivalence.map_homological_complex (e : V ≌ W) [e.functor.additive] (c : complex_shape ι): + homological_complex V c ≌ homological_complex W c := +{ functor := e.functor.map_homological_complex c, + inverse := e.inverse.map_homological_complex c, + unit_iso := (functor.map_homological_complex_id_iso V c).symm ≪≫ + nat_iso.map_homological_complex e.unit_iso c, + counit_iso := nat_iso.map_homological_complex e.counit_iso c ≪≫ + functor.map_homological_complex_id_iso W c, } + end category_theory namespace chain_complex diff --git a/src/category_theory/idempotents/homological_complex.lean b/src/category_theory/idempotents/homological_complex.lean index 5db952eacb7aa..ebc8a609b0b24 100644 --- a/src/category_theory/idempotents/homological_complex.lean +++ b/src/category_theory/idempotents/homological_complex.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Joël Riou -/ -import algebra.homology.homological_complex +import algebra.homology.additive import category_theory.idempotents.karoubi /-! @@ -14,6 +14,9 @@ This file contains simplifications lemmas for categories `karoubi (homological_complex C c)` and the construction of an equivalence of categories `karoubi (homological_complex C c) ≌ homological_complex (karoubi C) c`. +When the category `C` is idempotent complete, it is shown that +`homological_complex (karoubi C) c` is also idempotent complete. + -/ namespace category_theory @@ -203,6 +206,14 @@ def karoubi_cochain_complex_equivalence : karoubi (cochain_complex C α) ≌ cochain_complex (karoubi C) α := karoubi_homological_complex_equivalence C (complex_shape.up α) +instance [is_idempotent_complete C] : is_idempotent_complete (homological_complex C c) := +begin + rw [is_idempotent_complete_iff_of_equivalence + ((to_karoubi_equivalence C).map_homological_complex c), + ← is_idempotent_complete_iff_of_equivalence (karoubi_homological_complex_equivalence C c)], + apply_instance, +end + end idempotents end category_theory diff --git a/src/category_theory/idempotents/karoubi.lean b/src/category_theory/idempotents/karoubi.lean index ec19521b7a1b3..16fc6fd644f28 100644 --- a/src/category_theory/idempotents/karoubi.lean +++ b/src/category_theory/idempotents/karoubi.lean @@ -209,6 +209,14 @@ def to_karoubi_is_equivalence [is_idempotent_complete C] : is_equivalence (to_karoubi C) := equivalence.of_fully_faithfully_ess_surj (to_karoubi C) +/-- The equivalence `C ≅ karoubi C` when `C` is idempotent complete. -/ +def to_karoubi_equivalence [is_idempotent_complete C] : C ≌ karoubi C := +by { haveI := to_karoubi_is_equivalence C, exact functor.as_equivalence (to_karoubi C), } + +instance to_karoubi_equivalence_functor_additive + [preadditive C] [is_idempotent_complete C] : + (to_karoubi_equivalence C).functor.additive := (infer_instance : (to_karoubi C).additive) + namespace karoubi variables {C} From eb810cf549db839dadf13353dbe69bac55acdbbc Mon Sep 17 00:00:00 2001 From: AlexKontorovich <58564076+AlexKontorovich@users.noreply.github.com> Date: Wed, 29 Mar 2023 04:43:48 +0000 Subject: [PATCH 0720/1291] feat(measure_theory/group/fundamental_domain): integral_eq_tsum (#18132) This is `integral` analogues of existing `lintegral` lemmas. Also added `lintegral` versions of `sum_lintegral` and fixing naming consistency with primes vs no primes. (The prime goes with the `\inv`.) Co-authored-by: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> --- .../group/fundamental_domain.lean | 84 +++++++++++++------ 1 file changed, 57 insertions(+), 27 deletions(-) diff --git a/src/measure_theory/group/fundamental_domain.lean b/src/measure_theory/group/fundamental_domain.lean index 69a214a0ac2fc..689cacb93d810 100644 --- a/src/measure_theory/group/fundamental_domain.lean +++ b/src/measure_theory/group/fundamental_domain.lean @@ -201,18 +201,26 @@ h.sum_restrict_of_ac (refl _) ∫⁻ x, f x ∂μ = ∑' g : G, ∫⁻ x in g • s, f x ∂μ := h.lintegral_eq_tsum_of_ac (refl _) f -@[to_additive] lemma set_lintegral_eq_tsum' (h : is_fundamental_domain G s μ) (f : α → ℝ≥0∞) +@[to_additive] lemma lintegral_eq_tsum' (h : is_fundamental_domain G s μ) (f : α → ℝ≥0∞) : + ∫⁻ x, f x ∂μ = ∑' g : G, ∫⁻ x in s, f (g⁻¹ • x) ∂μ := +calc ∫⁻ x, f x ∂μ = ∑' g : G, ∫⁻ x in g • s, f x ∂μ : h.lintegral_eq_tsum f +... = ∑' g : G, ∫⁻ x in g⁻¹ • s, f x ∂μ : ((equiv.inv G).tsum_eq _).symm +... = ∑' g : G, ∫⁻ x in s, f (g⁻¹ • x) ∂μ : + tsum_congr $ λ g, ((measure_preserving_smul g⁻¹ μ).set_lintegral_comp_emb + (measurable_embedding_const_smul _) _ _).symm + +@[to_additive] lemma set_lintegral_eq_tsum (h : is_fundamental_domain G s μ) (f : α → ℝ≥0∞) (t : set α) : ∫⁻ x in t, f x ∂μ = ∑' g : G, ∫⁻ x in t ∩ g • s, f x ∂μ := calc ∫⁻ x in t, f x ∂μ = ∑' g : G, ∫⁻ x in g • s, f x ∂(μ.restrict t) : h.lintegral_eq_tsum_of_ac restrict_le_self.absolutely_continuous _ ... = ∑' g : G, ∫⁻ x in t ∩ g • s, f x ∂μ : by simp only [h.restrict_restrict, inter_comm] -@[to_additive] lemma set_lintegral_eq_tsum (h : is_fundamental_domain G s μ) (f : α → ℝ≥0∞) +@[to_additive] lemma set_lintegral_eq_tsum' (h : is_fundamental_domain G s μ) (f : α → ℝ≥0∞) (t : set α) : ∫⁻ x in t, f x ∂μ = ∑' g : G, ∫⁻ x in g • t ∩ s, f (g⁻¹ • x) ∂μ := calc ∫⁻ x in t, f x ∂μ = ∑' g : G, ∫⁻ x in t ∩ g • s, f x ∂μ : - h.set_lintegral_eq_tsum' f t + h.set_lintegral_eq_tsum f t ... = ∑' g : G, ∫⁻ x in t ∩ g⁻¹ • s, f x ∂μ : ((equiv.inv G).tsum_eq _).symm ... = ∑' g : G, ∫⁻ x in g⁻¹ • (g • t ∩ s), f (x) ∂μ : by simp only [smul_set_inter, inv_smul_smul] @@ -234,7 +242,7 @@ h.measure_eq_tsum_of_ac absolutely_continuous.rfl t @[to_additive] lemma measure_eq_tsum (h : is_fundamental_domain G s μ) (t : set α) : μ t = ∑' g : G, μ (g • t ∩ s) := -by simpa only [set_lintegral_one] using h.set_lintegral_eq_tsum (λ _, 1) t +by simpa only [set_lintegral_one] using h.set_lintegral_eq_tsum' (λ _, 1) t @[to_additive] lemma measure_zero_of_invariant (h : is_fundamental_domain G s μ) (t : set α) (ht : ∀ g : G, g • t = t) (hts : μ (t ∩ s) = 0) : @@ -261,9 +269,9 @@ end @[to_additive] protected lemma set_lintegral_eq (hs : is_fundamental_domain G s μ) (ht : is_fundamental_domain G t μ) (f : α → ℝ≥0∞) (hf : ∀ (g : G) x, f (g • x) = f x) : ∫⁻ x in s, f x ∂μ = ∫⁻ x in t, f x ∂μ := -calc ∫⁻ x in s, f x ∂μ = ∑' g : G, ∫⁻ x in s ∩ g • t, f x ∂μ : ht.set_lintegral_eq_tsum' _ _ +calc ∫⁻ x in s, f x ∂μ = ∑' g : G, ∫⁻ x in s ∩ g • t, f x ∂μ : ht.set_lintegral_eq_tsum _ _ ... = ∑' g : G, ∫⁻ x in g • t ∩ s, f (g⁻¹ • x) ∂μ : by simp only [hf, inter_comm] -... = ∫⁻ x in t, f x ∂μ : (hs.set_lintegral_eq_tsum _ _).symm +... = ∫⁻ x in t, f x ∂μ : (hs.set_lintegral_eq_tsum' _ _).symm @[to_additive] lemma measure_set_eq (hs : is_fundamental_domain G s μ) (ht : is_fundamental_domain G t μ) {A : set α} (hA₀ : measurable_set A) @@ -326,33 +334,55 @@ and_congr (hs.ae_strongly_measurable_on_iff ht hf) (hs.has_finite_integral_on_if variables [normed_space ℝ E] [complete_space E] +@[to_additive] lemma integral_eq_tsum_of_ac (h : is_fundamental_domain G s μ) (hν : ν ≪ μ) + (f : α → E) (hf : integrable f ν) : ∫ x, f x ∂ν = ∑' g : G, ∫ x in g • s, f x ∂ν := +begin + rw [← measure_theory.integral_sum_measure, h.sum_restrict_of_ac hν], + rw h.sum_restrict_of_ac hν, -- Weirdly, these rewrites seem not to be combinable + exact hf, +end + +@[to_additive] lemma integral_eq_tsum (h : is_fundamental_domain G s μ) + (f : α → E) (hf : integrable f μ) : ∫ x, f x ∂μ = ∑' g : G, ∫ x in g • s, f x ∂μ := +integral_eq_tsum_of_ac h (by refl) f hf + +@[to_additive] lemma integral_eq_tsum' (h : is_fundamental_domain G s μ) + (f : α → E) (hf : integrable f μ) : ∫ x, f x ∂μ = ∑' g : G, ∫ x in s, f (g⁻¹ • x) ∂μ := +calc ∫ x, f x ∂μ = ∑' g : G, ∫ x in g • s, f x ∂μ : h.integral_eq_tsum f hf +... = ∑' g : G, ∫ x in g⁻¹ • s, f x ∂μ : ((equiv.inv G).tsum_eq _).symm +... = ∑' g : G, ∫ x in s, f (g⁻¹ • x) ∂μ : + tsum_congr $ λ g, (measure_preserving_smul g⁻¹ μ).set_integral_image_emb + (measurable_embedding_const_smul _) _ _ + +@[to_additive] lemma set_integral_eq_tsum (h : is_fundamental_domain G s μ) {f : α → E} + {t : set α} (hf : integrable_on f t μ) : + ∫ x in t, f x ∂μ = ∑' g : G, ∫ x in t ∩ g • s, f x ∂μ := +calc ∫ x in t, f x ∂μ = ∑' g : G, ∫ x in g • s, f x ∂(μ.restrict t) : + h.integral_eq_tsum_of_ac restrict_le_self.absolutely_continuous f hf +... = ∑' g : G, ∫ x in t ∩ g • s, f x ∂μ : + by simp only [h.restrict_restrict, measure_smul, inter_comm] + +@[to_additive] lemma set_integral_eq_tsum' (h : is_fundamental_domain G s μ) {f : α → E} + {t : set α} (hf : integrable_on f t μ) : + ∫ x in t, f x ∂μ = ∑' g : G, ∫ x in g • t ∩ s, f (g⁻¹ • x) ∂μ := +calc ∫ x in t, f x ∂μ = ∑' g : G, ∫ x in t ∩ g • s, f x ∂μ : + h.set_integral_eq_tsum hf +... = ∑' g : G, ∫ x in t ∩ g⁻¹ • s, f x ∂μ : ((equiv.inv G).tsum_eq _).symm +... = ∑' g : G, ∫ x in g⁻¹ • (g • t ∩ s), f (x) ∂μ : + by simp only [smul_set_inter, inv_smul_smul] +... = ∑' g : G, ∫ x in g • t ∩ s, f (g⁻¹ • x) ∂μ : + tsum_congr $ λ g, (measure_preserving_smul g⁻¹ μ).set_integral_image_emb + (measurable_embedding_const_smul _) _ _ + @[to_additive] protected lemma set_integral_eq (hs : is_fundamental_domain G s μ) (ht : is_fundamental_domain G t μ) {f : α → E} (hf : ∀ (g : G) x, f (g • x) = f x) : ∫ x in s, f x ∂μ = ∫ x in t, f x ∂μ := begin by_cases hfs : integrable_on f s μ, { have hft : integrable_on f t μ, by rwa ht.integrable_on_iff hs hf, - have hac : ∀ {u}, μ.restrict u ≪ μ := λ u, restrict_le_self.absolutely_continuous, - calc ∫ x in s, f x ∂μ = ∫ x in ⋃ g : G, g • t, f x ∂(μ.restrict s) : - by rw [restrict_congr_set (hac ht.Union_smul_ae_eq), restrict_univ] - ... = ∑' g : G, ∫ x in g • t, f x ∂(μ.restrict s) : - integral_Union_ae (λ g, (ht.null_measurable_set_smul g).mono_ac hac) - (ht.pairwise_ae_disjoint_of_ac hac) hfs.integrable.integrable_on - ... = ∑' g : G, ∫ x in s ∩ g • t, f x ∂μ : - by simp only [ht.restrict_restrict, inter_comm] - ... = ∑' g : G, ∫ x in s ∩ g⁻¹ • t, f x ∂μ : ((equiv.inv G).tsum_eq _).symm - ... = ∑' g : G, ∫ x in g⁻¹ • (g • s ∩ t), f x ∂μ : - by simp only [smul_set_inter, inv_smul_smul] - ... = ∑' g : G, ∫ x in g • s ∩ t, f (g⁻¹ • x) ∂μ : - tsum_congr $ λ g, (measure_preserving_smul g⁻¹ μ).set_integral_image_emb - (measurable_embedding_const_smul _) _ _ - ... = ∑' g : G, ∫ x in g • s, f x ∂(μ.restrict t) : - by simp only [hf, hs.restrict_restrict] - ... = ∫ x in ⋃ g : G, g • s, f x ∂(μ.restrict t) : - (integral_Union_ae (λ g, (hs.null_measurable_set_smul g).mono_ac hac) - (hs.ae_disjoint.mono $ λ i j h, hac h) hft.integrable.integrable_on).symm - ... = ∫ x in t, f x ∂μ : - by rw [restrict_congr_set (hac hs.Union_smul_ae_eq), restrict_univ] }, + calc ∫ x in s, f x ∂μ = ∑' g : G, ∫ x in s ∩ g • t, f x ∂μ : ht.set_integral_eq_tsum hfs + ... = ∑' g : G, ∫ x in g • t ∩ s, f (g⁻¹ • x) ∂μ : by simp only [hf, inter_comm] + ... = ∫ x in t, f x ∂μ : (hs.set_integral_eq_tsum' hft).symm, }, { rw [integral_undef hfs, integral_undef], rwa [hs.integrable_on_iff ht hf] at hfs } end From bcbee715ab85a4f516c814effdf232618c0322af Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Mar=C3=ADa=20In=C3=A9s=20de=20Frutos-Fern=C3=A1ndez?= Date: Wed, 29 Mar 2023 06:00:28 +0000 Subject: [PATCH 0721/1291] feat(ring_theory/mv_polynomial/weighted homogeneous): add weighted homogeneous polynomials (#17855) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit It is possible to assign weights (in a commutative additive monoid `M`) to the variables of a multivariate polynomial ring, so that monomials of the ring then have a weighted degree with respect to the weights of the variables. The weights are represented by a function `w : σ → M`, where `σ` are the indeterminates. A multivariate polynomial `φ` is weighted homogeneous of weighted degree `m : M` if all monomials occuring in `φ` have the same weighted degree `m`. Co-authored-by: Antoine Chambert-Loir --- .../mv_polynomial/weighted_homogeneous.lean | 438 ++++++++++++++++++ 1 file changed, 438 insertions(+) create mode 100644 src/ring_theory/mv_polynomial/weighted_homogeneous.lean diff --git a/src/ring_theory/mv_polynomial/weighted_homogeneous.lean b/src/ring_theory/mv_polynomial/weighted_homogeneous.lean new file mode 100644 index 0000000000000..6d6edcc4d1f1c --- /dev/null +++ b/src/ring_theory/mv_polynomial/weighted_homogeneous.lean @@ -0,0 +1,438 @@ +/- +Copyright (c) 2022 María Inés de Frutos-Fernández. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Antoine Chambert-Loir, María Inés de Frutos-Fernández +-/ +import algebra.graded_monoid +import data.mv_polynomial.variables + +/-! +# Weighted homogeneous polynomials + +It is possible to assign weights (in a commutative additive monoid `M`) to the variables of a +multivariate polynomial ring, so that monomials of the ring then have a weighted degree with +respect to the weights of the variables. The weights are represented by a function `w : σ → M`, +where `σ` are the indeterminates. + +A multivariate polynomial `φ` is weighted homogeneous of weighted degree `m : M` if all monomials +occuring in `φ` have the same weighted degree `m`. + +## Main definitions/lemmas + +* `weighted_total_degree' w φ` : the weighted total degree of a multivariate polynomial with respect +to the weights `w`, taking values in `with_bot M`. + +* `weighted_total_degree w φ` : When `M` has a `⊥` element, we can define the weighted total degree +of a multivariate polynomial as a function taking values in `M`. + +* `is_weighted_homogeneous w φ m`: a predicate that asserts that `φ` is weighted homogeneous +of weighted degree `m` with respect to the weights `w`. + +* `weighted_homogeneous_submodule R w m`: the submodule of homogeneous polynomials +of weighted degree `m`. + +* `weighted_homogeneous_component w m`: the additive morphism that projects polynomials +onto their summand that is weighted homogeneous of degree `n` with respect to `w`. + +* `sum_weighted_homogeneous_component`: every polynomial is the sum of its weighted homogeneous +components. +-/ + +noncomputable theory + +open_locale classical big_operators + +open set function finset finsupp add_monoid_algebra + +variables {R M : Type*} [comm_semiring R] + +namespace mv_polynomial +variables {σ : Type*} + +section add_comm_monoid +variables [add_comm_monoid M] + +/-! ### `weighted_degree'` -/ + +/-- The `weighted degree'` of the finitely supported function `s : σ →₀ ℕ` is the sum + `∑(s i)•(w i)`. -/ +def weighted_degree' (w : σ → M) : (σ →₀ ℕ) →+ M := +(finsupp.total σ M ℕ w).to_add_monoid_hom + +section semilattice_sup +variable [semilattice_sup M] + +/-- The weighted total degree of a multivariate polynomial, taking values in `with_bot M`. -/ +def weighted_total_degree' (w : σ → M) (p : mv_polynomial σ R) : with_bot M := +p.support.sup (λ s, weighted_degree' w s) + +/-- The `weighted_total_degree'` of a polynomial `p` is `⊥` if and only if `p = 0`. -/ +lemma weighted_total_degree'_eq_bot_iff (w : σ → M) (p : mv_polynomial σ R) : + weighted_total_degree' w p = ⊥ ↔ p = 0 := +begin + simp only [weighted_total_degree',finset.sup_eq_bot_iff, mem_support_iff, with_bot.coe_ne_bot, + mv_polynomial.eq_zero_iff ], + exact forall_congr (λ _, not_not) +end + +/-- The `weighted_total_degree'` of the zero polynomial is `⊥`. -/ +lemma weighted_total_degree'_zero (w : σ → M) : + weighted_total_degree' w (0 : mv_polynomial σ R) = ⊥ := +by simp only [weighted_total_degree', support_zero, finset.sup_empty] + +section order_bot +variable [order_bot M] + +/-- When `M` has a `⊥` element, we can define the weighted total degree of a multivariate + polynomial as a function taking values in `M`. -/ +def weighted_total_degree (w : σ → M) (p : mv_polynomial σ R) : M := +p.support.sup (λ s, weighted_degree' w s) + +/-- This lemma relates `weighted_total_degree` and `weighted_total_degree'`. -/ +lemma weighted_total_degree_coe (w : σ → M) (p : mv_polynomial σ R) (hp : p ≠ 0): + weighted_total_degree' w p = ↑(weighted_total_degree w p) := +begin + rw [ne.def, ← weighted_total_degree'_eq_bot_iff w p, ← ne.def, with_bot.ne_bot_iff_exists] at hp, + obtain ⟨m, hm⟩ := hp, + apply le_antisymm, + { simp only [weighted_total_degree, weighted_total_degree', finset.sup_le_iff, + with_bot.coe_le_coe], + intro b, + exact finset.le_sup }, + { simp only [weighted_total_degree], + have hm' : weighted_total_degree' w p ≤ m := le_of_eq hm.symm, + rw ← hm, + simpa [weighted_total_degree'] using hm' } +end + +/-- The `weighted_total_degree` of the zero polynomial is `⊥`. -/ +lemma weighted_total_degree_zero (w : σ → M) : + weighted_total_degree w (0 : mv_polynomial σ R) = ⊥ := +by simp only [weighted_total_degree, support_zero, finset.sup_empty] + +lemma le_weighted_total_degree (w : σ → M) {φ : mv_polynomial σ R} {d : σ →₀ ℕ} + (hd : d ∈ φ.support) : weighted_degree' w d ≤ φ.weighted_total_degree w := +le_sup hd + +end order_bot +end semilattice_sup + +/-- A multivariate polynomial `φ` is weighted homogeneous of weighted degree `m` if all monomials + occuring in `φ` have weighted degree `m`. -/ +def is_weighted_homogeneous (w : σ → M) (φ : mv_polynomial σ R) (m : M) : Prop := +∀ ⦃d⦄, coeff d φ ≠ 0 → weighted_degree' w d = m + +variable (R) + +/-- The submodule of homogeneous `mv_polynomial`s of degree `n`. -/ +def weighted_homogeneous_submodule (w : σ → M) (m : M) : + submodule R (mv_polynomial σ R) := +{ carrier := { x | x.is_weighted_homogeneous w m }, + smul_mem' := λ r a ha c hc, begin + rw coeff_smul at hc, + exact ha (right_ne_zero_of_mul hc), + end, + zero_mem' := λ d hd, false.elim (hd $ coeff_zero _), + add_mem' := λ a b ha hb c hc, begin + rw coeff_add at hc, + obtain h|h : coeff c a ≠ 0 ∨ coeff c b ≠ 0, + { contrapose! hc, simp only [hc, add_zero] }, + { exact ha h }, + { exact hb h }, + end } + +@[simp] lemma mem_weighted_homogeneous_submodule (w : σ → M) (m : M) (p : mv_polynomial σ R) : + p ∈ weighted_homogeneous_submodule R w m ↔ p.is_weighted_homogeneous w m := iff.rfl + +variables (R) + +/-- The submodule ` weighted_homogeneous_submodule R w m` of homogeneous `mv_polynomial`s of + degree `n` is equal to the `R`-submodule of all `p : (σ →₀ ℕ) →₀ R` such that + `p.support ⊆ {d | weighted_degree' w d = m}`. While equal, the former has a + convenient definitional reduction. -/ +lemma weighted_homogeneous_submodule_eq_finsupp_supported (w : σ → M) (m : M) : + weighted_homogeneous_submodule R w m = + finsupp.supported _ R {d | weighted_degree' w d = m} := +begin + ext, + simp only [mem_supported, set.subset_def, finsupp.mem_support_iff, mem_coe], + refl, +end + +variables {R} + +/-- The submodule generated by products `Pm *Pn` of weighted homogeneous polynomials of degrees `m` + and `n` is contained in the submodule of weighted homogeneous polynomials of degree `m + n`. -/ +lemma weighted_homogeneous_submodule_mul (w : σ → M) (m n : M) : + weighted_homogeneous_submodule R w m * weighted_homogeneous_submodule R w n ≤ + weighted_homogeneous_submodule R w (m + n) := +begin + rw submodule.mul_le, + intros φ hφ ψ hψ c hc, + rw [coeff_mul] at hc, + obtain ⟨⟨d, e⟩, hde, H⟩ := finset.exists_ne_zero_of_sum_ne_zero hc, + have aux : coeff d φ ≠ 0 ∧ coeff e ψ ≠ 0, + { contrapose! H, + by_cases h : coeff d φ = 0; + simp only [*, ne.def, not_false_iff, zero_mul, mul_zero] at * }, + rw [← (finsupp.mem_antidiagonal.mp hde), ← hφ aux.1, ← hψ aux.2, map_add], +end + +/-- Monomials are weighted homogeneous. -/ +lemma is_weighted_homogeneous_monomial (w : σ → M) (d : σ →₀ ℕ) (r : R) {m : M} + (hm : weighted_degree' w d = m) : is_weighted_homogeneous w (monomial d r) m := +begin + intros c hc, + rw coeff_monomial at hc, + split_ifs at hc with h, + { subst c, exact hm }, + { contradiction } +end + +/-- A polynomial of weighted_total_degree `⊥` is weighted_homogeneous of degree `⊥`. -/ +lemma is_weighted_homogeneous_of_total_degree_zero [semilattice_sup M] [order_bot M] + (w : σ → M) {p : mv_polynomial σ R} (hp : weighted_total_degree w p = (⊥ : M)) : + is_weighted_homogeneous w p (⊥ : M) := +begin + intros d hd, + have h := weighted_total_degree_coe w p (mv_polynomial.ne_zero_iff.mpr ⟨d, hd⟩), + simp only [weighted_total_degree', hp] at h, + rw [eq_bot_iff, ← with_bot.coe_le_coe, ← h], + exact finset.le_sup (mem_support_iff.mpr hd), +end + +/-- Constant polynomials are weighted homogeneous of degree 0. -/ +lemma is_weighted_homogeneous_C (w : σ → M) (r : R) : + is_weighted_homogeneous w (C r : mv_polynomial σ R) 0 := +is_weighted_homogeneous_monomial _ _ _ (map_zero _) + +variables (R) + +/-- 0 is weighted homogeneous of any degree. -/ +lemma is_weighted_homogeneous_zero (w : σ → M) (m : M) : + is_weighted_homogeneous w (0 : mv_polynomial σ R) m := +(weighted_homogeneous_submodule R w m).zero_mem + +/-- 1 is weighted homogeneous of degree 0. -/ +lemma is_weighted_homogeneous_one (w : σ → M) : + is_weighted_homogeneous w (1 : mv_polynomial σ R) 0 := +is_weighted_homogeneous_C _ _ + +/-- An indeterminate `i : σ` is weighted homogeneous of degree `w i`. -/ +lemma is_weighted_homogeneous_X (w : σ → M) (i : σ) : + is_weighted_homogeneous w (X i : mv_polynomial σ R) (w i) := +begin + apply is_weighted_homogeneous_monomial, + simp only [weighted_degree', linear_map.to_add_monoid_hom_coe, total_single, one_nsmul], +end + +namespace is_weighted_homogeneous +variables {R} {φ ψ : mv_polynomial σ R} {m n : M} + +/-- The weighted degree of a weighted homogeneous polynomial controls its support. -/ +lemma coeff_eq_zero {w : σ → M} (hφ : is_weighted_homogeneous w φ n) (d : σ →₀ ℕ) + (hd : weighted_degree' w d ≠ n) : coeff d φ = 0 := +by { have aux := mt (@hφ d) hd, rwa not_not at aux } + +/-- The weighted degree of a nonzero weighted homogeneous polynomial is well-defined. -/ +lemma inj_right {w : σ → M} (hφ : φ ≠ 0) (hm : is_weighted_homogeneous w φ m) + (hn : is_weighted_homogeneous w φ n) : m = n := +begin + obtain ⟨d, hd⟩ : ∃ d, coeff d φ ≠ 0 := exists_coeff_ne_zero hφ, + rw [← hm hd, ← hn hd] +end + +/-- The sum of two weighted homogeneous polynomials of degree `n` is weighted homogeneous of + weighted degree `n`. -/ +lemma add {w : σ → M} (hφ : is_weighted_homogeneous w φ n) (hψ : is_weighted_homogeneous w ψ n) : + is_weighted_homogeneous w (φ + ψ) n := +(weighted_homogeneous_submodule R w n).add_mem hφ hψ + +/-- The sum of weighted homogeneous polynomials of degree `n` is weighted homogeneous of + weighted degree `n`. -/ +lemma sum {ι : Type*} (s : finset ι) (φ : ι → mv_polynomial σ R) (n : M) {w : σ → M} + (h : ∀ i ∈ s, is_weighted_homogeneous w (φ i) n) : + is_weighted_homogeneous w (∑ i in s, φ i) n := +(weighted_homogeneous_submodule R w n).sum_mem h + +/-- The product of weighted homogeneous polynomials of weighted degrees `m` and `n` is weighted + homogeneous of weighted degree `m + n`. -/ +lemma mul {w : σ → M} (hφ : is_weighted_homogeneous w φ m) (hψ : is_weighted_homogeneous w ψ n) : + is_weighted_homogeneous w (φ * ψ) (m + n) := +weighted_homogeneous_submodule_mul w m n $ submodule.mul_mem_mul hφ hψ + +/-- A product of weighted homogeneous polynomials is weighted homogeneous, with weighted degree + equal to the sum of the weighted degrees. -/ +lemma prod {ι : Type*} (s : finset ι) (φ : ι → mv_polynomial σ R) (n : ι → M) {w : σ → M} : + (∀ i ∈ s, is_weighted_homogeneous w (φ i) (n i)) → + is_weighted_homogeneous w (∏ i in s, φ i) (∑ i in s, n i) := +begin + apply finset.induction_on s, + { intro, simp only [is_weighted_homogeneous_one, finset.sum_empty, finset.prod_empty] }, + { intros i s his IH h, + simp only [his, finset.prod_insert, finset.sum_insert, not_false_iff], + apply (h i (finset.mem_insert_self _ _)).mul (IH _), + intros j hjs, + exact h j (finset.mem_insert_of_mem hjs) } +end + +/-- A non zero weighted homogeneous polynomial of weighted degree `n` has weighted total degree + `n`. -/ +lemma weighted_total_degree [semilattice_sup M] {w : σ → M} (hφ : is_weighted_homogeneous w φ n) + (h : φ ≠ 0) : weighted_total_degree' w φ = n := +begin + simp only [weighted_total_degree'], + apply le_antisymm, + { simp only [finset.sup_le_iff, mem_support_iff, with_bot.coe_le_coe], + exact λ d hd, le_of_eq (hφ hd), }, + { obtain ⟨d, hd⟩ : ∃ d, coeff d φ ≠ 0 := exists_coeff_ne_zero h, + simp only [← hφ hd, finsupp.sum], + replace hd := finsupp.mem_support_iff.mpr hd, + exact finset.le_sup hd, } +end + +/-- The weighted homogeneous submodules form a graded monoid. -/ +instance weighted_homogeneous_submodule.gcomm_monoid {w : σ → M} : + set_like.graded_monoid (weighted_homogeneous_submodule R w) := +{ one_mem := is_weighted_homogeneous_one R w, + mul_mem := λ i j xi xj, is_weighted_homogeneous.mul } + +end is_weighted_homogeneous + +variables {R} + +/-- `weighted_homogeneous_component w n φ` is the part of `φ` that is weighted homogeneous of + weighted degree `n`, with respect to the weights `w`. + See `sum_weighted_homogeneous_component` for the statement that `φ` is equal to the sum + of all its weighted homogeneous components. -/ +def weighted_homogeneous_component (w : σ → M) (n : M) : + mv_polynomial σ R →ₗ[R] mv_polynomial σ R := +(submodule.subtype _).comp $ finsupp.restrict_dom _ _ {d | weighted_degree' w d = n} + +section weighted_homogeneous_component + +variables {w : σ → M} (n : M) (φ ψ : mv_polynomial σ R) + +lemma coeff_weighted_homogeneous_component (d : σ →₀ ℕ) : + coeff d (weighted_homogeneous_component w n φ) = + if weighted_degree' w d = n then coeff d φ else 0 := +finsupp.filter_apply (λ d : σ →₀ ℕ, weighted_degree' w d = n) φ d + +lemma weighted_homogeneous_component_apply : + weighted_homogeneous_component w n φ = + ∑ d in φ.support.filter (λ d, weighted_degree' w d = n), monomial d (coeff d φ) := +finsupp.filter_eq_sum (λ d : σ →₀ ℕ, weighted_degree' w d = n) φ + +/-- The `n` weighted homogeneous component of a polynomial is weighted homogeneous of +weighted degree `n`. -/ +lemma weighted_homogeneous_component_is_weighted_homogeneous : + (weighted_homogeneous_component w n φ).is_weighted_homogeneous w n := +begin + intros d hd, + contrapose! hd, + rw [coeff_weighted_homogeneous_component, if_neg hd] +end + +@[simp] lemma weighted_homogeneous_component_C_mul (n : M) (r : R) : + weighted_homogeneous_component w n (C r * φ) = C r * weighted_homogeneous_component w n φ := +by simp only [C_mul', linear_map.map_smul] + +lemma weighted_homogeneous_component_eq_zero' (h : ∀ d : σ →₀ ℕ, d ∈ φ.support → + weighted_degree' w d ≠ n) : weighted_homogeneous_component w n φ = 0 := +begin + rw [weighted_homogeneous_component_apply, sum_eq_zero], + intros d hd, rw mem_filter at hd, + exfalso, exact h _ hd.1 hd.2 +end + +lemma weighted_homogeneous_component_eq_zero [semilattice_sup M] [order_bot M] + (h : weighted_total_degree w φ < n) : weighted_homogeneous_component w n φ = 0 := +begin + rw [weighted_homogeneous_component_apply, sum_eq_zero], + intros d hd, rw mem_filter at hd, + exfalso, + apply lt_irrefl n, + nth_rewrite 0 ← hd.2, + exact lt_of_le_of_lt (le_weighted_total_degree w hd.1) h, +end + +lemma weighted_homogeneous_component_finsupp : + (function.support (λ m, weighted_homogeneous_component w m φ)).finite := +begin + suffices : function.support (λ m, weighted_homogeneous_component w m φ) ⊆ + (λ d, weighted_degree' w d) '' φ.support, + { exact finite.subset ((λ (d : σ →₀ ℕ), (weighted_degree' w) d) '' ↑(support φ)).to_finite this }, + intros m hm, + by_contradiction hm', apply hm, + simp only [mem_support, ne.def] at hm, + simp only [set.mem_image, not_exists, not_and] at hm', + exact weighted_homogeneous_component_eq_zero' m φ hm', +end + +variable (w) + +/-- Every polynomial is the sum of its weighted homogeneous components. -/ +lemma sum_weighted_homogeneous_component : + finsum (λ m, weighted_homogeneous_component w m φ) = φ := +begin + rw finsum_eq_sum _ (weighted_homogeneous_component_finsupp φ), + ext1 d, + simp only [coeff_sum, coeff_weighted_homogeneous_component], + rw finset.sum_eq_single (weighted_degree' w d), + { rw if_pos rfl, }, + { intros m hm hm', rw if_neg hm'.symm, }, + { intro hm, rw if_pos rfl, + simp only [finite.mem_to_finset, mem_support, ne.def, not_not] at hm, + have := coeff_weighted_homogeneous_component _ φ d, + rw [hm, if_pos rfl, coeff_zero] at this, + exact this.symm, }, +end + +variable {w} + +/-- The weighted homogeneous components of a weighted homogeneous polynomial. -/ +lemma weighted_homogeneous_component_weighted_homogeneous_polynomial (m n : M) + (p : mv_polynomial σ R) (h : p ∈ weighted_homogeneous_submodule R w n) : + weighted_homogeneous_component w m p = if m = n then p else 0 := +begin + simp only [mem_weighted_homogeneous_submodule] at h, + ext x, + rw coeff_weighted_homogeneous_component, + by_cases zero_coeff : coeff x p = 0, + { split_ifs, + all_goals { simp only [zero_coeff, coeff_zero], }, }, + { rw h zero_coeff, + simp only [show n = m ↔ m = n, from eq_comm], + split_ifs with h1, + { refl }, + { simp only [coeff_zero] } } +end + +end weighted_homogeneous_component + +end add_comm_monoid + +section canonically_ordered_add_monoid + +variables [canonically_ordered_add_monoid M] {w : σ → M} (φ : mv_polynomial σ R) + +/-- If `M` is a `canonically_ordered_add_monoid`, then the `weighted_homogeneous_component` + of weighted degree `0` of a polynomial is its constant coefficient. -/ +@[simp] lemma weighted_homogeneous_component_zero [no_zero_smul_divisors ℕ M] + (hw : ∀ i : σ, w i ≠ 0) : weighted_homogeneous_component w 0 φ = C (coeff 0 φ) := +begin + ext1 d, + rcases em (d = 0) with (rfl|hd), + { simp only [coeff_weighted_homogeneous_component, if_pos, map_zero, coeff_zero_C] }, + { rw [coeff_weighted_homogeneous_component, if_neg, coeff_C, if_neg (ne.symm hd)], + simp only [weighted_degree', linear_map.to_add_monoid_hom_coe, finsupp.total_apply, + finsupp.sum, sum_eq_zero_iff, finsupp.mem_support_iff, ne.def, smul_eq_zero, + not_forall, not_or_distrib, and_self_left, exists_prop], + simp only [finsupp.ext_iff, finsupp.coe_zero, pi.zero_apply, not_forall] at hd, + obtain ⟨i, hi⟩ := hd, + exact ⟨i, hi, hw i⟩ } +end + +end canonically_ordered_add_monoid + +end mv_polynomial From 2de9c37fa71dde2f1c6feff19876dd6a7b1519f0 Mon Sep 17 00:00:00 2001 From: Moritz Doll Date: Wed, 29 Mar 2023 07:17:29 +0000 Subject: [PATCH 0722/1291] chore(linear_algebra/affine_space/combination): make k explicit in affine_combination (#18689) The implicitness caused problems in elaboration. In Lean 3 it only amounts to long elaboration times, but in Lean 4 elaboration fails. --- src/analysis/convex/between.lean | 14 +--- src/analysis/convex/combination.lean | 6 +- src/geometry/euclidean/basic.lean | 4 +- src/geometry/euclidean/circumcenter.lean | 8 +- src/geometry/euclidean/monge_point.lean | 2 +- src/linear_algebra/affine_space/basis.lean | 10 +-- .../affine_space/combination.lean | 78 ++++++++++--------- .../affine_space/independent.lean | 23 +++--- src/linear_algebra/affine_space/matrix.lean | 2 +- 9 files changed, 73 insertions(+), 74 deletions(-) diff --git a/src/analysis/convex/between.lean b/src/analysis/convex/between.lean index 1c1e2653347de..5d192618f9e2c 100644 --- a/src/analysis/convex/between.lean +++ b/src/analysis/convex/between.lean @@ -455,20 +455,14 @@ lemma sbtw.trans_wbtw_right_ne [no_zero_smul_divisors R V] {w x y z : P} (h₁ : (h₂ : wbtw R x y z) : w ≠ y := h₁.wbtw.trans_right_ne h₂ h₁.left_ne -/- Calls to `affine_combination` are slow to elaborate (generally, not just for this lemma), and -without the use of `@finset.affine_combination R V _ _ _ _ _ _` for at least three of the six -calls in this lemma statement, elaboration of the statement times out (even if the proof is -replaced by `sorry`). -/ lemma sbtw.affine_combination_of_mem_affine_span_pair [no_zero_divisors R] [no_zero_smul_divisors R V] {ι : Type*} {p : ι → P} (ha : affine_independent R p) {w w₁ w₂ : ι → R} {s : finset ι} (hw : ∑ i in s, w i = 1) (hw₁ : ∑ i in s, w₁ i = 1) (hw₂ : ∑ i in s, w₂ i = 1) - (h : s.affine_combination p w ∈ - line[R, s.affine_combination p w₁, s.affine_combination p w₂]) {i : ι} (his : i ∈ s) - (hs : sbtw R (w₁ i) (w i) (w₂ i)) : - sbtw R (@finset.affine_combination R V _ _ _ _ _ _ s p w₁) - (@finset.affine_combination R V _ _ _ _ _ _ s p w) - (@finset.affine_combination R V _ _ _ _ _ _ s p w₂) := + (h : s.affine_combination R p w ∈ + line[R, s.affine_combination R p w₁, s.affine_combination R p w₂]) + {i : ι} (his : i ∈ s) (hs : sbtw R (w₁ i) (w i) (w₂ i)) : + sbtw R (s.affine_combination R p w₁) (s.affine_combination R p w) (s.affine_combination R p w₂) := begin rw affine_combination_mem_affine_span_pair ha hw hw₁ hw₂ at h, rcases h with ⟨r, hr⟩, diff --git a/src/analysis/convex/combination.lean b/src/analysis/convex/combination.lean index f21d458d49b99..ee1a808317e50 100644 --- a/src/analysis/convex/combination.lean +++ b/src/analysis/convex/combination.lean @@ -223,7 +223,7 @@ t.center_mass_mem_convex_hull hw₀ hws (λ i, mem_coe.2) lemma affine_combination_eq_center_mass {ι : Type*} {t : finset ι} {p : ι → E} {w : ι → R} (hw₂ : ∑ i in t, w i = 1) : - affine_combination t p w = center_mass t w p := + t.affine_combination R p w = center_mass t w p := begin rw [affine_combination_eq_weighted_vsub_of_point_vadd_of_sum_eq_one _ w _ hw₂ (0 : E), finset.weighted_vsub_of_point_apply, vadd_eq_add, add_zero, t.center_mass_eq_of_sum_1 _ hw₂], @@ -232,7 +232,7 @@ end lemma affine_combination_mem_convex_hull {s : finset ι} {v : ι → E} {w : ι → R} (hw₀ : ∀ i ∈ s, 0 ≤ w i) (hw₁ : s.sum w = 1) : - s.affine_combination v w ∈ convex_hull R (range v) := + s.affine_combination R v w ∈ convex_hull R (range v) := begin rw affine_combination_eq_center_mass hw₁, apply s.center_mass_mem_convex_hull hw₀, @@ -258,7 +258,7 @@ end lemma convex_hull_range_eq_exists_affine_combination (v : ι → E) : convex_hull R (range v) = { x | ∃ (s : finset ι) (w : ι → R) - (hw₀ : ∀ i ∈ s, 0 ≤ w i) (hw₁ : s.sum w = 1), s.affine_combination v w = x } := + (hw₀ : ∀ i ∈ s, 0 ≤ w i) (hw₁ : s.sum w = 1), s.affine_combination R v w = x } := begin refine subset.antisymm (convex_hull_min _ _) _, { intros x hx, diff --git a/src/geometry/euclidean/basic.lean b/src/geometry/euclidean/basic.lean index 86f03e5a75ac3..543b4470efd8d 100644 --- a/src/geometry/euclidean/basic.lean +++ b/src/geometry/euclidean/basic.lean @@ -86,11 +86,11 @@ in terms of the pairwise distances between the points in that combination. -/ lemma dist_affine_combination {ι : Type*} {s : finset ι} {w₁ w₂ : ι → ℝ} (p : ι → P) (h₁ : ∑ i in s, w₁ i = 1) (h₂ : ∑ i in s, w₂ i = 1) : - by have a₁ := s.affine_combination p w₁; have a₂ := s.affine_combination p w₂; exact + by have a₁ := s.affine_combination ℝ p w₁; have a₂ := s.affine_combination ℝ p w₂; exact dist a₁ a₂ * dist a₁ a₂ = (-∑ i₁ in s, ∑ i₂ in s, (w₁ - w₂) i₁ * (w₁ - w₂) i₂ * (dist (p i₁) (p i₂) * dist (p i₁) (p i₂))) / 2 := begin - rw [dist_eq_norm_vsub V (s.affine_combination p w₁) (s.affine_combination p w₂), + rw [dist_eq_norm_vsub V (s.affine_combination ℝ p w₁) (s.affine_combination ℝ p w₂), ←@inner_self_eq_norm_mul_norm ℝ, finset.affine_combination_vsub], have h : ∑ i in s, (w₁ - w₂) i = 0, { simp_rw [pi.sub_apply, finset.sum_sub_distrib, h₁, h₂, sub_self] }, diff --git a/src/geometry/euclidean/circumcenter.lean b/src/geometry/euclidean/circumcenter.lean index e25e36dc17fa7..41736c1778e66 100644 --- a/src/geometry/euclidean/circumcenter.lean +++ b/src/geometry/euclidean/circumcenter.lean @@ -583,7 +583,7 @@ include V lemma point_eq_affine_combination_of_points_with_circumcenter {n : ℕ} (s : simplex ℝ P n) (i : fin (n + 1)) : s.points i = - (univ : finset (points_with_circumcenter_index n)).affine_combination + (univ : finset (points_with_circumcenter_index n)).affine_combination ℝ s.points_with_circumcenter (point_weights_with_circumcenter i) := begin rw ←points_with_circumcenter_point, @@ -627,7 +627,7 @@ include V lemma centroid_eq_affine_combination_of_points_with_circumcenter {n : ℕ} (s : simplex ℝ P n) (fs : finset (fin (n + 1))) : fs.centroid ℝ s.points = - (univ : finset (points_with_circumcenter_index n)).affine_combination + (univ : finset (points_with_circumcenter_index n)).affine_combination ℝ s.points_with_circumcenter (centroid_weights_with_circumcenter fs) := begin simp_rw [centroid_def, affine_combination_apply, @@ -666,7 +666,7 @@ include V `points_with_circumcenter`. -/ lemma circumcenter_eq_affine_combination_of_points_with_circumcenter {n : ℕ} (s : simplex ℝ P n) : - s.circumcenter = (univ : finset (points_with_circumcenter_index n)).affine_combination + s.circumcenter = (univ : finset (points_with_circumcenter_index n)).affine_combination ℝ s.points_with_circumcenter (circumcenter_weights_with_circumcenter n) := begin rw ←points_with_circumcenter_eq_circumcenter, @@ -702,7 +702,7 @@ terms of `points_with_circumcenter`. -/ lemma reflection_circumcenter_eq_affine_combination_of_points_with_circumcenter {n : ℕ} (s : simplex ℝ P n) {i₁ i₂ : fin (n + 1)} (h : i₁ ≠ i₂) : reflection (affine_span ℝ (s.points '' {i₁, i₂})) s.circumcenter = - (univ : finset (points_with_circumcenter_index n)).affine_combination + (univ : finset (points_with_circumcenter_index n)).affine_combination ℝ s.points_with_circumcenter (reflection_circumcenter_weights_with_circumcenter i₁ i₂) := begin have hc : card ({i₁, i₂} : finset (fin (n + 1))) = 2, diff --git a/src/geometry/euclidean/monge_point.lean b/src/geometry/euclidean/monge_point.lean index 981a00e57a85e..f6eb1d8e00df1 100644 --- a/src/geometry/euclidean/monge_point.lean +++ b/src/geometry/euclidean/monge_point.lean @@ -126,7 +126,7 @@ include V `points_with_circumcenter`. -/ lemma monge_point_eq_affine_combination_of_points_with_circumcenter {n : ℕ} (s : simplex ℝ P (n + 2)) : - s.monge_point = (univ : finset (points_with_circumcenter_index (n + 2))).affine_combination + s.monge_point = (univ : finset (points_with_circumcenter_index (n + 2))).affine_combination ℝ s.points_with_circumcenter (monge_point_weights_with_circumcenter n) := begin rw [monge_point_eq_smul_vsub_vadd_circumcenter, diff --git a/src/linear_algebra/affine_space/basis.lean b/src/linear_algebra/affine_space/basis.lean index 8cf3de70cf3de..42e4192d7ea65 100644 --- a/src/linear_algebra/affine_space/basis.lean +++ b/src/linear_algebra/affine_space/basis.lean @@ -145,7 +145,7 @@ lemma coord_apply [decidable_eq ι] (i j : ι) : b.coord i (b j) = if i = j then by cases eq_or_ne i j; simp [h] @[simp] lemma coord_apply_combination_of_mem (hi : i ∈ s) {w : ι → k} (hw : s.sum w = 1) : - b.coord i (s.affine_combination b w) = w i := + b.coord i (s.affine_combination k b w) = w i := begin classical, simp only [coord_apply, hi, finset.affine_combination_eq_linear_combination, if_true, mul_boole, @@ -153,7 +153,7 @@ begin end @[simp] lemma coord_apply_combination_of_not_mem (hi : i ∉ s) {w : ι → k} (hw : s.sum w = 1) : - b.coord i (s.affine_combination b w) = 0 := + b.coord i (s.affine_combination k b w) = 0 := begin classical, simp only [coord_apply, hi, finset.affine_combination_eq_linear_combination, if_false, mul_boole, @@ -171,7 +171,7 @@ begin end @[simp] lemma affine_combination_coord_eq_self [fintype ι] (q : P) : - finset.univ.affine_combination b (λ i, b.coord i q) = q := + finset.univ.affine_combination k b (λ i, b.coord i q) = q := begin have hq : q ∈ affine_span k (range b), { rw b.tot, exact affine_subspace.mem_top k V q, }, obtain ⟨w, hw, rfl⟩ := eq_affine_combination_of_mem_affine_span_of_fintype hq, @@ -208,7 +208,7 @@ begin let s : finset ι := {i}, have hi : i ∈ s, { simp, }, have hw : s.sum (function.const ι (1 : k)) = 1, { simp, }, - have hq : q = s.affine_combination b (function.const ι (1 : k)), { simp, }, + have hq : q = s.affine_combination k b (function.const ι (1 : k)), { simp, }, rw [pi.one_apply, hq, b.coord_apply_combination_of_mem hi hw], end @@ -223,7 +223,7 @@ begin have hj : j ∈ s, { simp, }, let w : ι → k := λ j', if j' = i then x else 1-x, have hw : s.sum w = 1, { simp [hij, finset.sum_ite, finset.filter_insert, finset.filter_eq'], }, - use s.affine_combination b w, + use s.affine_combination k b w, simp [b.coord_apply_combination_of_mem hi hw], end diff --git a/src/linear_algebra/affine_space/combination.lean b/src/linear_algebra/affine_space/combination.lean index d2bd478803367..41c118950e3ff 100644 --- a/src/linear_algebra/affine_space/combination.lean +++ b/src/linear_algebra/affine_space/combination.lean @@ -331,6 +331,8 @@ lemma weighted_vsub_const_smul (w : ι → k) (p : ι → P) (c : k) : s.weighted_vsub p (c • w) = c • s.weighted_vsub p w := s.weighted_vsub_of_point_const_smul _ _ _ _ +variables (k) + /-- A weighted sum of the results of subtracting a default base point from the given points, added to that base point, as an affine map on the weights. This is intended to be used when the sum of the weights @@ -346,9 +348,11 @@ def affine_combination (p : ι → P) : (ι → k) →ᵃ[k] P := /-- The linear map corresponding to `affine_combination` is `weighted_vsub`. -/ @[simp] lemma affine_combination_linear (p : ι → P) : - (s.affine_combination p : (ι → k) →ᵃ[k] P).linear = s.weighted_vsub p := + (s.affine_combination k p).linear = s.weighted_vsub p := rfl +variables {k} + /-- Applying `affine_combination` with given weights. This is for the case where a result involving a default base point is OK (for example, when that base point will cancel out later); a more typical use case @@ -357,41 +361,41 @@ point with `affine_combination_eq_weighted_vsub_of_point_vadd_of_sum_eq_one` and then using `weighted_vsub_of_point_apply`. -/ lemma affine_combination_apply (w : ι → k) (p : ι → P) : - s.affine_combination p w = + s.affine_combination k p w = s.weighted_vsub_of_point p (classical.choice S.nonempty) w +ᵥ (classical.choice S.nonempty) := rfl /-- The value of `affine_combination`, where the given points are equal. -/ @[simp] lemma affine_combination_apply_const (w : ι → k) (p : P) (h : ∑ i in s, w i = 1) : - s.affine_combination (λ _, p) w = p := + s.affine_combination k (λ _, p) w = p := by rw [affine_combination_apply, s.weighted_vsub_of_point_apply_const, h, one_smul, vsub_vadd] /-- `affine_combination` gives equal results for two families of weights and two families of points that are equal on `s`. -/ lemma affine_combination_congr {w₁ w₂ : ι → k} (hw : ∀ i ∈ s, w₁ i = w₂ i) {p₁ p₂ : ι → P} - (hp : ∀ i ∈ s, p₁ i = p₂ i) : s.affine_combination p₁ w₁ = s.affine_combination p₂ w₂ := + (hp : ∀ i ∈ s, p₁ i = p₂ i) : s.affine_combination k p₁ w₁ = s.affine_combination k p₂ w₂ := by simp_rw [affine_combination_apply, s.weighted_vsub_of_point_congr hw hp] /-- `affine_combination` gives the sum with any base point, when the sum of the weights is 1. -/ lemma affine_combination_eq_weighted_vsub_of_point_vadd_of_sum_eq_one (w : ι → k) (p : ι → P) (h : ∑ i in s, w i = 1) (b : P) : - s.affine_combination p w = s.weighted_vsub_of_point p b w +ᵥ b := + s.affine_combination k p w = s.weighted_vsub_of_point p b w +ᵥ b := s.weighted_vsub_of_point_vadd_eq_of_sum_eq_one w p h _ _ /-- Adding a `weighted_vsub` to an `affine_combination`. -/ lemma weighted_vsub_vadd_affine_combination (w₁ w₂ : ι → k) (p : ι → P) : - s.weighted_vsub p w₁ +ᵥ s.affine_combination p w₂ = s.affine_combination p (w₁ + w₂) := + s.weighted_vsub p w₁ +ᵥ s.affine_combination k p w₂ = s.affine_combination k p (w₁ + w₂) := by rw [←vadd_eq_add, affine_map.map_vadd, affine_combination_linear] /-- Subtracting two `affine_combination`s. -/ lemma affine_combination_vsub (w₁ w₂ : ι → k) (p : ι → P) : - s.affine_combination p w₁ -ᵥ s.affine_combination p w₂ = s.weighted_vsub p (w₁ - w₂) := + s.affine_combination k p w₁ -ᵥ s.affine_combination k p w₂ = s.weighted_vsub p (w₁ - w₂) := by rw [←affine_map.linear_map_vsub, affine_combination_linear, vsub_eq_sub] lemma attach_affine_combination_of_injective [decidable_eq P] (s : finset P) (w : P → k) (f : s → P) (hf : function.injective f) : - s.attach.affine_combination f (w ∘ f) = (image f univ).affine_combination id w := + s.attach.affine_combination k f (w ∘ f) = (image f univ).affine_combination k id w := begin simp only [affine_combination, weighted_vsub_of_point_apply, id.def, vadd_right_cancel_iff, function.comp_app, affine_map.coe_mk], @@ -404,7 +408,7 @@ begin end lemma attach_affine_combination_coe (s : finset P) (w : P → k) : - s.attach.affine_combination (coe : s → P) (w ∘ coe) = s.affine_combination id w := + s.attach.affine_combination k (coe : s → P) (w ∘ coe) = s.affine_combination k id w := by classical; rw [attach_affine_combination_of_injective s w (coe : s → P) subtype.coe_injective, univ_eq_attach, attach_image_coe] @@ -421,7 +425,7 @@ by simp [s.weighted_vsub_apply, vsub_eq_sub, smul_sub, ← finset.sum_smul, hw] combinations. -/ @[simp] lemma affine_combination_eq_linear_combination (s : finset ι) (p : ι → V) (w : ι → k) (hw : ∑ i in s, w i = 1) : - s.affine_combination p w = ∑ i in s, w i • p i := + s.affine_combination k p w = ∑ i in s, w i • p i := by simp [s.affine_combination_eq_weighted_vsub_of_point_vadd_of_sum_eq_one w p hw 0] include S @@ -430,7 +434,7 @@ include S and has weight 1 and the other points in the set have weight 0. -/ @[simp] lemma affine_combination_of_eq_one_of_eq_zero (w : ι → k) (p : ι → P) {i : ι} (his : i ∈ s) (hwi : w i = 1) (hw0 : ∀ i2 ∈ s, i2 ≠ i → w i2 = 0) : - s.affine_combination p w = p i := + s.affine_combination k p w = p i := begin have h1 : ∑ i in s, w i = 1 := hwi ▸ sum_eq_single i hw0 (λ h, false.elim (h his)), rw [s.affine_combination_eq_weighted_vsub_of_point_vadd_of_sum_eq_one w p h1 (p i), @@ -447,7 +451,7 @@ end corresponding indicator function and adding points to the set. -/ lemma affine_combination_indicator_subset (w : ι → k) (p : ι → P) {s₁ s₂ : finset ι} (h : s₁ ⊆ s₂) : - s₁.affine_combination p w = s₂.affine_combination p (set.indicator ↑s₁ w) := + s₁.affine_combination k p w = s₂.affine_combination k p (set.indicator ↑s₁ w) := by rw [affine_combination_apply, affine_combination_apply, weighted_vsub_of_point_indicator_subset _ _ _ h] @@ -455,13 +459,13 @@ by rw [affine_combination_apply, affine_combination_apply, affine combination with the same points and weights over the original `finset`. -/ lemma affine_combination_map (e : ι₂ ↪ ι) (w : ι → k) (p : ι → P) : - (s₂.map e).affine_combination p w = s₂.affine_combination (p ∘ e) (w ∘ e) := + (s₂.map e).affine_combination k p w = s₂.affine_combination k (p ∘ e) (w ∘ e) := by simp_rw [affine_combination_apply, weighted_vsub_of_point_map] /-- A weighted sum of pairwise subtractions, expressed as a subtraction of two `affine_combination` expressions. -/ lemma sum_smul_vsub_eq_affine_combination_vsub (w : ι → k) (p₁ p₂ : ι → P) : - ∑ i in s, w i • (p₁ i -ᵥ p₂ i) = s.affine_combination p₁ w -ᵥ s.affine_combination p₂ w := + ∑ i in s, w i • (p₁ i -ᵥ p₂ i) = s.affine_combination k p₁ w -ᵥ s.affine_combination k p₂ w := begin simp_rw [affine_combination_apply, vadd_vsub_vadd_cancel_right], exact s.sum_smul_vsub_eq_weighted_vsub_of_point_sub _ _ _ _ @@ -471,20 +475,20 @@ end sum of the weights is 1. -/ lemma sum_smul_vsub_const_eq_affine_combination_vsub (w : ι → k) (p₁ : ι → P) (p₂ : P) (h : ∑ i in s, w i = 1) : - ∑ i in s, w i • (p₁ i -ᵥ p₂) = s.affine_combination p₁ w -ᵥ p₂ := + ∑ i in s, w i • (p₁ i -ᵥ p₂) = s.affine_combination k p₁ w -ᵥ p₂ := by rw [sum_smul_vsub_eq_affine_combination_vsub, affine_combination_apply_const _ _ _ h] /-- A weighted sum of pairwise subtractions, where the point on the left is constant and the sum of the weights is 1. -/ lemma sum_smul_const_vsub_eq_vsub_affine_combination (w : ι → k) (p₂ : ι → P) (p₁ : P) (h : ∑ i in s, w i = 1) : - ∑ i in s, w i • (p₁ -ᵥ p₂ i) = p₁ -ᵥ s.affine_combination p₂ w := + ∑ i in s, w i • (p₁ -ᵥ p₂ i) = p₁ -ᵥ s.affine_combination k p₂ w := by rw [sum_smul_vsub_eq_affine_combination_vsub, affine_combination_apply_const _ _ _ h] /-- A weighted sum may be split into a subtraction of affine combinations over two subsets. -/ lemma affine_combination_sdiff_sub [decidable_eq ι] {s₂ : finset ι} (h : s₂ ⊆ s) (w : ι → k) (p : ι → P) : - (s \ s₂).affine_combination p w -ᵥ s₂.affine_combination p (-w) = s.weighted_vsub p w := + (s \ s₂).affine_combination k p w -ᵥ s₂.affine_combination k p (-w) = s.weighted_vsub p w := begin simp_rw [affine_combination_apply, vadd_vsub_vadd_cancel_right], exact s.weighted_vsub_sdiff_sub h _ _ @@ -494,7 +498,7 @@ end the affine combination of the other points with the given weights. -/ lemma affine_combination_eq_of_weighted_vsub_eq_zero_of_eq_neg_one {w : ι → k} {p : ι → P} (hw : s.weighted_vsub p w = (0 : V)) {i : ι} [decidable_pred (≠ i)] (his : i ∈ s) - (hwi : w i = -1) : (s.filter (≠ i)).affine_combination p w = p i := + (hwi : w i = -1) : (s.filter (≠ i)).affine_combination k p w = p i := begin classical, rw [←@vsub_eq_zero_iff_eq V, ←hw, @@ -509,15 +513,15 @@ end /-- An affine combination over `s.subtype pred` equals one over `s.filter pred`. -/ lemma affine_combination_subtype_eq_filter (w : ι → k) (p : ι → P) (pred : ι → Prop) [decidable_pred pred] : - (s.subtype pred).affine_combination (λ i, p i) (λ i, w i) = - (s.filter pred).affine_combination p w := + (s.subtype pred).affine_combination k (λ i, p i) (λ i, w i) = + (s.filter pred).affine_combination k p w := by rw [affine_combination_apply, affine_combination_apply, weighted_vsub_of_point_subtype_eq_filter] /-- An affine combination over `s.filter pred` equals one over `s` if all the weights at indices in `s` not satisfying `pred` are zero. -/ lemma affine_combination_filter_of_ne (w : ι → k) (p : ι → P) {pred : ι → Prop} [decidable_pred pred] (h : ∀ i ∈ s, w i ≠ 0 → pred i) : - (s.filter pred).affine_combination p w = s.affine_combination p w := + (s.filter pred).affine_combination k p w = s.affine_combination k p w := by rw [affine_combination_apply, affine_combination_apply, s.weighted_vsub_of_point_filter_of_ne _ _ _ h] @@ -575,9 +579,9 @@ subset. -/ lemma eq_affine_combination_subset_iff_eq_affine_combination_subtype {p0 : P} {s : set ι} {p : ι → P} : (∃ (fs : finset ι) (hfs : ↑fs ⊆ s) (w : ι → k) (hw : ∑ i in fs, w i = 1), - p0 = fs.affine_combination p w) ↔ + p0 = fs.affine_combination k p w) ↔ ∃ (fs : finset s) (w : s → k) (hw : ∑ i in fs, w i = 1), - p0 = fs.affine_combination (λ (i : s), p i) w := + p0 = fs.affine_combination k (λ (i : s), p i) w := begin simp_rw [affine_combination_apply, eq_vadd_iff_vsub_eq], exact eq_weighted_vsub_of_point_subset_iff_eq_weighted_vsub_of_point_subtype @@ -588,7 +592,7 @@ variables {k V} /-- Affine maps commute with affine combinations. -/ lemma map_affine_combination {V₂ P₂ : Type*} [add_comm_group V₂] [module k V₂] [affine_space V₂ P₂] (p : ι → P) (w : ι → k) (hw : s.sum w = 1) (f : P →ᵃ[k] P₂) : - f (s.affine_combination p w) = s.affine_combination (f ∘ p) w := + f (s.affine_combination k p w) = s.affine_combination k (f ∘ p) w := begin have b := classical.choice (infer_instance : affine_space V P).nonempty, have b₂ := classical.choice (infer_instance : affine_space V₂ P₂).nonempty, @@ -684,7 +688,7 @@ variables (k) /-- An affine combination with `affine_combination_single_weights` gives the specified point. -/ @[simp] lemma affine_combination_affine_combination_single_weights [decidable_eq ι] (p : ι → P) - {i : ι} (hi : i ∈ s) : s.affine_combination p (affine_combination_single_weights k i) = p i := + {i : ι} (hi : i ∈ s) : s.affine_combination k p (affine_combination_single_weights k i) = p i := begin refine s.affine_combination_of_eq_one_of_eq_zero _ _ hi (by simp) _, rintro j - hj, @@ -707,7 +711,7 @@ variables {k} `line_map`. -/ @[simp] lemma affine_combination_affine_combination_line_map_weights [decidable_eq ι] (p : ι → P) {i j : ι} (hi : i ∈ s) (hj : j ∈ s) (c : k) : - s.affine_combination p (affine_combination_line_map_weights i j c) = + s.affine_combination k p (affine_combination_line_map_weights i j c) = affine_map.line_map (p i) (p j) c := by rw [affine_combination_line_map_weights, ←weighted_vsub_vadd_affine_combination, weighted_vsub_const_smul, s.affine_combination_affine_combination_single_weights k p hi, @@ -766,11 +770,11 @@ include V is intended to be used in the case where the number of points, converted to `k`, is not zero. -/ def centroid (p : ι → P) : P := -s.affine_combination p (s.centroid_weights k) +s.affine_combination k p (s.centroid_weights k) /-- The definition of the centroid. -/ lemma centroid_def (p : ι → P) : - s.centroid k p = s.affine_combination p (s.centroid_weights k) := + s.centroid k p = s.affine_combination k p (s.centroid_weights k) := rfl lemma centroid_univ (s : finset P) : @@ -865,7 +869,7 @@ include V /-- The centroid as an affine combination over a `fintype`. -/ lemma centroid_eq_affine_combination_fintype [fintype ι] (p : ι → P) : - s.centroid k p = univ.affine_combination p (s.centroid_weights_indicator k) := + s.centroid k p = univ.affine_combination k p (s.centroid_weights_indicator k) := affine_combination_indicator_subset _ _ (subset_univ _) /-- An indexed family of points that is injective on the given @@ -948,7 +952,7 @@ end nontrivial. -/ lemma affine_combination_mem_affine_span [nontrivial k] {s : finset ι} {w : ι → k} (h : ∑ i in s, w i = 1) (p : ι → P) : - s.affine_combination p w ∈ affine_span k (set.range p) := + s.affine_combination k p w ∈ affine_span k (set.range p) := begin classical, have hnz : ∑ i in s, w i ≠ 0 := h.symm ▸ one_ne_zero, @@ -957,14 +961,14 @@ begin let w1 : ι → k := function.update (function.const ι 0) i1 1, have hw1 : ∑ i in s, w1 i = 1, { rw [finset.sum_update_of_mem hi1, finset.sum_const_zero, add_zero] }, - have hw1s : s.affine_combination p w1 = p i1 := + have hw1s : s.affine_combination k p w1 = p i1 := s.affine_combination_of_eq_one_of_eq_zero w1 p hi1 (function.update_same _ _ _) (λ _ _ hne, function.update_noteq hne _ _), - have hv : s.affine_combination p w -ᵥ p i1 ∈ (affine_span k (set.range p)).direction, + have hv : s.affine_combination k p w -ᵥ p i1 ∈ (affine_span k (set.range p)).direction, { rw [direction_affine_span, ←hw1s, finset.affine_combination_vsub], apply weighted_vsub_mem_vector_span, simp [pi.sub_apply, h, hw1] }, - rw ←vsub_vadd (s.affine_combination p w) (p i1), + rw ←vsub_vadd (s.affine_combination k p w) (p i1), exact affine_subspace.vadd_mem_of_mem_direction hv (mem_affine_span k (set.mem_range_self _)) end @@ -1019,7 +1023,7 @@ variables {k} `eq_affine_combination_of_mem_affine_span_of_fintype`. -/ lemma eq_affine_combination_of_mem_affine_span {p1 : P} {p : ι → P} (h : p1 ∈ affine_span k (set.range p)) : - ∃ (s : finset ι) (w : ι → k) (hw : ∑ i in s, w i = 1), p1 = s.affine_combination p w := + ∃ (s : finset ι) (w : ι → k) (hw : ∑ i in s, w i = 1), p1 = s.affine_combination k p w := begin classical, have hn : ((affine_span k (set.range p)) : set P).nonempty := ⟨p1, h⟩, @@ -1040,7 +1044,7 @@ begin let w0 : ι → k := function.update (function.const ι 0) i0 1, have hw0 : ∑ i in s', w0 i = 1, { rw [finset.sum_update_of_mem (finset.mem_insert_self _ _), finset.sum_const_zero, add_zero] }, - have hw0s : s'.affine_combination p w0 = p i0 := + have hw0s : s'.affine_combination k p w0 = p i0 := s'.affine_combination_of_eq_one_of_eq_zero w0 p (finset.mem_insert_self _ _) (function.update_same _ _ _) @@ -1053,7 +1057,7 @@ end lemma eq_affine_combination_of_mem_affine_span_of_fintype [fintype ι] {p1 : P} {p : ι → P} (h : p1 ∈ affine_span k (set.range p)) : - ∃ (w : ι → k) (hw : ∑ i, w i = 1), p1 = finset.univ.affine_combination p w := + ∃ (w : ι → k) (hw : ∑ i, w i = 1), p1 = finset.univ.affine_combination k p w := begin classical, obtain ⟨s, w, hw, rfl⟩ := eq_affine_combination_of_mem_affine_span h, @@ -1069,7 +1073,7 @@ if it is an `affine_combination` with sum of weights 1, provided the underlying ring is nontrivial. -/ lemma mem_affine_span_iff_eq_affine_combination [nontrivial k] {p1 : P} {p : ι → P} : p1 ∈ affine_span k (set.range p) ↔ - ∃ (s : finset ι) (w : ι → k) (hw : ∑ i in s, w i = 1), p1 = s.affine_combination p w := + ∃ (s : finset ι) (w : ι → k) (hw : ∑ i in s, w i = 1), p1 = s.affine_combination k p w := begin split, { exact eq_affine_combination_of_mem_affine_span }, diff --git a/src/linear_algebra/affine_space/independent.lean b/src/linear_algebra/affine_space/independent.lean index 1bd97d4143948..e781e031dfef7 100644 --- a/src/linear_algebra/affine_space/independent.lean +++ b/src/linear_algebra/affine_space/independent.lean @@ -176,7 +176,7 @@ combinations (with sum of weights 1) that evaluate to the same point have equal `set.indicator`. -/ lemma affine_independent_iff_indicator_eq_of_affine_combination_eq (p : ι → P) : affine_independent k p ↔ ∀ (s1 s2 : finset ι) (w1 w2 : ι → k), ∑ i in s1, w1 i = 1 → - ∑ i in s2, w2 i = 1 → s1.affine_combination p w1 = s2.affine_combination p w2 → + ∑ i in s2, w2 i = 1 → s1.affine_combination k p w1 = s2.affine_combination k p w2 → set.indicator ↑s1 w1 = set.indicator ↑s2 w2 := begin classical, @@ -199,13 +199,13 @@ begin let w1 : ι → k := function.update (function.const ι 0) i0 1, have hw1 : ∑ i in s, w1 i = 1, { rw [finset.sum_update_of_mem hi0, finset.sum_const_zero, add_zero] }, - have hw1s : s.affine_combination p w1 = p i0 := + have hw1s : s.affine_combination k p w1 = p i0 := s.affine_combination_of_eq_one_of_eq_zero w1 p hi0 (function.update_same _ _ _) (λ _ _ hne, function.update_noteq hne _ _), let w2 := w + w1, have hw2 : ∑ i in s, w2 i = 1, { simp [w2, finset.sum_add_distrib, hw, hw1] }, - have hw2s : s.affine_combination p w2 = p i0, + have hw2s : s.affine_combination k p w2 = p i0, { simp [w2, ←finset.weighted_vsub_vadd_affine_combination, hs, hw1s] }, replace ha := ha s s w2 w1 hw2 hw1 (hw1s.symm ▸ hw2s), have hws : w2 i0 - w1 i0 = 0, @@ -218,7 +218,7 @@ end combinations (with sum of weights 1) that evaluate to the same point are equal. -/ lemma affine_independent_iff_eq_of_fintype_affine_combination_eq [fintype ι] (p : ι → P) : affine_independent k p ↔ ∀ (w1 w2 : ι → k), ∑ i, w1 i = 1 → ∑ i, w2 i = 1 → - finset.univ.affine_combination p w1 = finset.univ.affine_combination p w2 → w1 = w2 := + finset.univ.affine_combination k p w1 = finset.univ.affine_combination k p w2 → w1 = w2 := begin rw affine_independent_iff_indicator_eq_of_affine_combination_eq, split, @@ -252,7 +252,7 @@ end lemma affine_independent.indicator_eq_of_affine_combination_eq {p : ι → P} (ha : affine_independent k p) (s₁ s₂ : finset ι) (w₁ w₂ : ι → k) (hw₁ : ∑ i in s₁, w₁ i = 1) - (hw₂ : ∑ i in s₂, w₂ i = 1) (h : s₁.affine_combination p w₁ = s₂.affine_combination p w₂) : + (hw₂ : ∑ i in s₂, w₂ i = 1) (h : s₁.affine_combination k p w₁ = s₂.affine_combination k p w₂) : set.indicator ↑s₁ w₁ = set.indicator ↑s₂ w₂ := (affine_independent_iff_indicator_eq_of_affine_combination_eq k p).1 ha s₁ s₂ w₁ w₂ hw₁ hw₂ h @@ -479,7 +479,7 @@ lemma weighted_vsub_mem_vector_span_pair {p : ι → P} (h : affine_independent {w w₁ w₂ : ι → k} {s : finset ι} (hw : ∑ i in s, w i = 0) (hw₁ : ∑ i in s, w₁ i = 1) (hw₂ : ∑ i in s, w₂ i = 1) : s.weighted_vsub p w ∈ - vector_span k ({s.affine_combination p w₁, s.affine_combination p w₂} : set P) ↔ + vector_span k ({s.affine_combination k p w₁, s.affine_combination k p w₂} : set P) ↔ ∃ r : k, ∀ i ∈ s, w i = r * (w₁ i - w₂ i) := begin rw mem_vector_span_pair, @@ -509,11 +509,11 @@ two points. -/ lemma affine_combination_mem_affine_span_pair {p : ι → P} (h : affine_independent k p) {w w₁ w₂ : ι → k} {s : finset ι} (hw : ∑ i in s, w i = 1) (hw₁ : ∑ i in s, w₁ i = 1) (hw₂ : ∑ i in s, w₂ i = 1) : - s.affine_combination p w ∈ - line[k, s.affine_combination p w₁, s.affine_combination p w₂] ↔ + s.affine_combination k p w ∈ + line[k, s.affine_combination k p w₁, s.affine_combination k p w₂] ↔ ∃ r : k, ∀ i ∈ s, w i = r * (w₂ i - w₁ i) + w₁ i := begin - rw [←vsub_vadd (s.affine_combination p w) (s.affine_combination p w₁), + rw [←vsub_vadd (s.affine_combination k p w) (s.affine_combination k p w₁), affine_subspace.vadd_mem_iff_mem_direction _ (left_mem_affine_span_pair _ _ _), direction_affine_span, s.affine_combination_vsub, set.pair_comm, weighted_vsub_mem_vector_span_pair h _ hw₂ hw₁], @@ -719,7 +719,8 @@ sign. -/ lemma sign_eq_of_affine_combination_mem_affine_span_pair {p : ι → P} (h : affine_independent k p) {w w₁ w₂ : ι → k} {s : finset ι} (hw : ∑ i in s, w i = 1) (hw₁ : ∑ i in s, w₁ i = 1) (hw₂ : ∑ i in s, w₂ i = 1) - (hs : s.affine_combination p w ∈ line[k, s.affine_combination p w₁, s.affine_combination p w₂]) + (hs : s.affine_combination k p w ∈ + line[k, s.affine_combination k p w₁, s.affine_combination k p w₂]) {i j : ι} (hi : i ∈ s) (hj : j ∈ s) (hi0 : w₁ i = 0) (hj0 : w₁ j = 0) (hij : sign (w₂ i) = sign (w₂ j)) : sign (w i) = sign (w j) := begin @@ -737,7 +738,7 @@ lemma sign_eq_of_affine_combination_mem_affine_span_single_line_map {p : ι → (h : affine_independent k p) {w : ι → k} {s : finset ι} (hw : ∑ i in s, w i = 1) {i₁ i₂ i₃ : ι} (h₁ : i₁ ∈ s) (h₂ : i₂ ∈ s) (h₃ : i₃ ∈ s) (h₁₂ : i₁ ≠ i₂) (h₁₃ : i₁ ≠ i₃) (h₂₃ : i₂ ≠ i₃) {c : k} (hc0 : 0 < c) (hc1 : c < 1) - (hs : s.affine_combination p w ∈ line[k, p i₁, affine_map.line_map (p i₂) (p i₃) c]) : + (hs : s.affine_combination k p w ∈ line[k, p i₁, affine_map.line_map (p i₂) (p i₃) c]) : sign (w i₂) = sign (w i₃) := begin classical, diff --git a/src/linear_algebra/affine_space/matrix.lean b/src/linear_algebra/affine_space/matrix.lean index 6376fea840f17..799e8203837cc 100644 --- a/src/linear_algebra/affine_space/matrix.lean +++ b/src/linear_algebra/affine_space/matrix.lean @@ -85,7 +85,7 @@ begin ... = ∑ l, ∑ j, (A i j) * b.to_matrix p j l : by rw finset.sum_comm ... = ∑ l, (A ⬝ b.to_matrix p) i l : rfl ... = 1 : by simp [hA, matrix.one_apply, finset.filter_eq], }, - have hbi : b i = finset.univ.affine_combination p (A i), + have hbi : b i = finset.univ.affine_combination k p (A i), { apply b.ext_elem, intros j, rw [b.coord_apply, finset.univ.map_affine_combination _ _ hAi, From 9cb7206107eb5cbc8dd3b9fc0864c548fae962bd Mon Sep 17 00:00:00 2001 From: PatrickKinnear Date: Wed, 29 Mar 2023 08:50:38 +0000 Subject: [PATCH 0723/1291] doc(representation_theory/basic): add docstring to dual_tensor_hom_comm (#18661) add a docstring to dual_tensor_hom_comm to describe its mathematical content explicitly --- src/representation_theory/basic.lean | 7 +++++++ 1 file changed, 7 insertions(+) diff --git a/src/representation_theory/basic.lean b/src/representation_theory/basic.lean index 2db82fe660fdd..81292da3bc65b 100644 --- a/src/representation_theory/basic.lean +++ b/src/representation_theory/basic.lean @@ -378,6 +378,13 @@ def dual : representation k G (module.dual k V) := @[simp] lemma dual_apply (g : G) : (dual ρV) g = module.dual.transpose (ρV g⁻¹) := rfl +/-- +Given $k$-modules $V, W$, there is a homomorphism $φ : V^* ⊗ W → Hom_k(V, W)$ +(implemented by `linear_algebra.contraction.dual_tensor_hom`). +Given representations of $G$ on $V$ and $W$,there are representations of $G$ on $V^* ⊗ W$ and on +$Hom_k(V, W)$. +This lemma says that $φ$ is $G$-linear. +-/ lemma dual_tensor_hom_comm (g : G) : (dual_tensor_hom k V W) ∘ₗ (tensor_product.map (ρV.dual g) (ρW g)) = (lin_hom ρV ρW) g ∘ₗ (dual_tensor_hom k V W) := From b685f506164f8d17a6404048bc4d696739c5d976 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 29 Mar 2023 11:40:01 +0000 Subject: [PATCH 0724/1291] =?UTF-8?q?feat(data/finset/pointwise):=20`a=20?= =?UTF-8?q?=E2=80=A2=20(s=20=E2=88=A9=20t)=20=3D=20a=20=E2=80=A2=20s=20?= =?UTF-8?q?=E2=88=A9=20a=20=E2=80=A2=20t`=20(#18682)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit I added the corresponding `set` lemmas a while back, but forgot about their `finset` counterpart. Also additivise `finset.is_central_scalar`/`set.is_central_scalar`. --- src/data/finset/image.lean | 14 +++++++++----- src/data/finset/pointwise.lean | 26 ++++++++++++++++++++++++++ src/data/set/pointwise/smul.lean | 1 + 3 files changed, 36 insertions(+), 5 deletions(-) diff --git a/src/data/finset/image.lean b/src/data/finset/image.lean index 31afc2348852a..e41fa0925240e 100644 --- a/src/data/finset/image.lean +++ b/src/data/finset/image.lean @@ -354,11 +354,7 @@ subset_inter (image_subset_image $ inter_subset_left _ _) $ lemma image_inter_of_inj_on [decidable_eq α] {f : α → β} (s t : finset α) (hf : set.inj_on f (s ∪ t)) : (s ∩ t).image f = s.image f ∩ t.image f := -(image_inter_subset _ _ _).antisymm $ λ x, begin - simp only [mem_inter, mem_image], - rintro ⟨⟨a, ha, rfl⟩, b, hb, h⟩, - exact ⟨a, ⟨ha, by rwa ←hf (or.inr hb) (or.inl ha) h⟩, rfl⟩, -end +coe_injective $ by { push_cast, exact set.image_inter_on (λ a ha b hb, hf (or.inr ha) $ or.inl hb) } lemma image_inter [decidable_eq α] (s₁ s₂ : finset α) (hf : injective f) : (s₁ ∩ s₂).image f = s₁.image f ∩ s₂.image f := @@ -392,6 +388,14 @@ end ⟨λ h, eq_empty_of_forall_not_mem $ λ a m, ne_empty_of_mem (mem_image_of_mem _ m) h, λ e, e.symm ▸ rfl⟩ +lemma image_sdiff [decidable_eq α] {f : α → β} (s t : finset α) (hf : injective f) : + (s \ t).image f = s.image f \ t.image f := +coe_injective $ by { push_cast, exact set.image_diff hf _ _ } + +lemma image_symm_diff [decidable_eq α] {f : α → β} (s t : finset α) (hf : injective f) : + (s ∆ t).image f = s.image f ∆ t.image f := +coe_injective $ by { push_cast, exact set.image_symm_diff hf _ _ } + @[simp] lemma _root_.disjoint.of_image_finset {s t : finset α} {f : α → β} (h : disjoint (s.image f) (t.image f)) : disjoint s t := diff --git a/src/data/finset/pointwise.lean b/src/data/finset/pointwise.lean index ac4a45def0c2b..d4833b70a33cb 100644 --- a/src/data/finset/pointwise.lean +++ b/src/data/finset/pointwise.lean @@ -876,6 +876,7 @@ instance is_scalar_tower'' [has_smul α β] [has_smul α γ] [has_smul β γ] [i is_scalar_tower (finset α) (finset β) (finset γ) := ⟨λ a s t, coe_injective $ by simp only [coe_smul_finset, coe_smul, smul_assoc]⟩ +@[to_additive] instance is_central_scalar [has_smul α β] [has_smul αᵐᵒᵖ β] [is_central_scalar α β] : is_central_scalar α (finset β) := ⟨λ a s, coe_injective $ by simp only [coe_smul_finset, coe_smul, op_smul_eq_smul]⟩ @@ -1002,6 +1003,22 @@ by { simp_rw ←coe_subset, push_cast, exact set.set_smul_subset_iff } @[to_additive] lemma subset_smul_finset_iff : s ⊆ a • t ↔ a⁻¹ • s ⊆ t := by { simp_rw ←coe_subset, push_cast, exact set.subset_set_smul_iff } +@[to_additive] lemma smul_finset_inter : a • (s ∩ t) = a • s ∩ a • t := +image_inter _ _ $ mul_action.injective a + +@[to_additive] lemma smul_finset_sdiff : a • (s \ t) = a • s \ a • t := +image_sdiff _ _ $ mul_action.injective a + +@[to_additive] lemma smul_finset_symm_diff : a • (s ∆ t) = (a • s) ∆ (a • t) := +image_symm_diff _ _ $ mul_action.injective a + +@[simp, to_additive] lemma smul_finset_univ [fintype β] : a • (univ : finset β) = univ := +image_univ_of_surjective $ mul_action.surjective a + +@[simp, to_additive] lemma smul_univ [fintype β] {s : finset α} (hs : s.nonempty) : + s • (univ : finset β) = univ := +coe_injective $ by { push_cast, exact set.smul_univ hs } + @[simp, to_additive] lemma card_smul_finset (a : α) (s : finset β) : (a • s).card = s.card := card_image_of_injective _ $ mul_action.injective _ @@ -1028,6 +1045,15 @@ show units.mk0 a ha • _ ⊆ _ ↔ _, from smul_finset_subset_iff lemma subset_smul_finset_iff₀ (ha : a ≠ 0) : s ⊆ a • t ↔ a⁻¹ • s ⊆ t := show _ ⊆ units.mk0 a ha • _ ↔ _, from subset_smul_finset_iff +lemma smul_finset_inter₀ (ha : a ≠ 0) : a • (s ∩ t) = a • s ∩ a • t := +image_inter _ _ $ mul_action.injective₀ ha + +lemma smul_finset_sdiff₀ (ha : a ≠ 0) : a • (s \ t) = a • s \ a • t := +image_sdiff _ _ $ mul_action.injective₀ ha + +lemma smul_finset_symm_diff₀ (ha : a ≠ 0) : a • (s ∆ t) = (a • s) ∆ (a • t) := +image_symm_diff _ _ $ mul_action.injective₀ ha + lemma smul_univ₀ [fintype β] {s : finset α} (hs : ¬ s ⊆ 0) : s • (univ : finset β) = univ := coe_injective $ by { rw ←coe_subset at hs, push_cast at ⊢ hs, exact set.smul_univ₀ hs } diff --git a/src/data/set/pointwise/smul.lean b/src/data/set/pointwise/smul.lean index 6900611e8e936..75e8c8d504d26 100644 --- a/src/data/set/pointwise/smul.lean +++ b/src/data/set/pointwise/smul.lean @@ -243,6 +243,7 @@ instance is_scalar_tower'' [has_smul α β] [has_smul α γ] [has_smul β γ] [i is_scalar_tower (set α) (set β) (set γ) := { smul_assoc := λ T T' T'', image2_assoc smul_assoc } +@[to_additive] instance is_central_scalar [has_smul α β] [has_smul αᵐᵒᵖ β] [is_central_scalar α β] : is_central_scalar α (set β) := ⟨λ a S, congr_arg (λ f, f '' S) $ by exact funext (λ _, op_smul_eq_smul _ _)⟩ From 1c775cc661988d96c477aa4ca6f7b5641a2a924b Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Thu, 30 Mar 2023 00:32:05 +0000 Subject: [PATCH 0725/1291] docs(tutorials/representation_theory): beginnings of a tutorial following Etingof's notes (#13911) Co-authored-by: Scott Morrison --- .../representation_theory/etingof.lean | 190 ++++++++++++++++++ src/algebra/category/Module/algebra.lean | 2 +- src/linear_algebra/matrix/to_lin.lean | 2 +- 3 files changed, 192 insertions(+), 2 deletions(-) create mode 100644 docs/tutorial/representation_theory/etingof.lean diff --git a/docs/tutorial/representation_theory/etingof.lean b/docs/tutorial/representation_theory/etingof.lean new file mode 100644 index 0000000000000..f56927ac0a8e2 --- /dev/null +++ b/docs/tutorial/representation_theory/etingof.lean @@ -0,0 +1,190 @@ +/- +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.simple +import category_theory.subobject.basic +import category_theory.preadditive.schur +import algebra.algebra.restrict_scalars +import algebra.algebra.tower +import algebra.category.Module.algebra +import algebra.category.Module.images +import algebra.category.Module.biproducts +import algebra.category.Module.simple +import data.mv_polynomial.basic +import algebra.free_algebra +import data.complex.module + +/-! +# "Introduction to representation theory" by Etingof + +This tutorial file follows along with the lecture notes "Introduction to representation theory", +by Pavel Etingof and other contributors. + +These lecture notes are available freely online at . + +This tutorial is (extremely) incomplete at present. +The goal is to work through the lecture notes, +showing how to use the definitions and results from mathlib +to establish the results in Etingof's notes. (We are not giving separate proofs here!) + +Our intention is (sadly) to skip all the problems, and many of the examples. + +Often results are proved by reference to (much) more general facts in mathlib. +-/ + +axiom skipped {p : Sort*} : p + +universes u +open category_theory finite_dimensional + +noncomputable theory + +/-! +## Chapter 2 "Basic notions of representation theory" +-/ + +/-! +### 2.2 "Algebras" +-/ + +-- Definition 2.2.1: An associative algebra. +variables {k : Type*} [field k] +variables {A : Type*} [ring A] [algebra k A] + +-- Etingof uses the word "unit" to refer to the identity in an algebra. +-- Currently in mathlib all algebras are unital +-- (although non-unital rings exists as `non_unital_ring`) +-- Thus we skip Definition 2.2.2 and Proposition 2.2.3 + +-- Example 2.2.4 (1)-(5) +example : algebra k k := by apply_instance +example {X : Type*} [fintype X] : algebra k (mv_polynomial X k) := by apply_instance +example {V : Type*} [add_comm_group V] [module k V] : algebra k (V →ₗ[k] V) := by apply_instance +example {X : Type*} : algebra k (free_algebra k X) := by apply_instance +example {G : Type*} [group G] : algebra k (monoid_algebra k G) := by apply_instance + +-- Definition 2.2.5: A commutative algebra is described as: +example {A : Type*} [comm_ring A] [algebra k A] := true + +-- Definition 2.2.6: algebra homomorphisms: +example {B : Type*} [ring B] [algebra k B] (f : A →ₐ[k] B) := true + +/-! +## 2.3 "Representations" +-/ + +-- Definition 2.3.1 +-- A representation (of an associative algebra) will usually be described as a module. +variables {M : Type*} [add_comm_group M] [module k M] [module A M] [is_scalar_tower k A M] + +-- or we can use `Module A`, for a "bundled" `A`-module, +-- which is useful when we want access to the category theory library. +variables (N : Module.{u} A) + +-- We can translate between these easily: +-- "bundle" a type with appropriate typeclasses +example : Module A := Module.of A M +-- a "bundled" module has a coercion to `Type`, +-- that comes equipped with the appropriate typeclasses: +example : module A N := by apply_instance + +-- Remark 2.3.2: Right `A`-modules are handled as left `Aᵐᵒᵖ`-modules: +example : Module Aᵐᵒᵖ := Module.of Aᵐᵒᵖ A +-- Right modules are not extensively developed in mathlib at this point, +-- and you may run into difficulty using them. + +-- It is helpful when working with `Module` to run +open_locale Module +-- which adds some instances. + +-- Example 2.3.3 +-- (1) The zero module +example : Module A := Module.of A punit +-- (2) The left regular module +example : Module A := Module.of A A +-- (3) Modules over a field are vector spaces. +-- (Because we handle vector spaces as modules in mathlib, this is empty of content.) +example (V : Type*) [add_comm_group V] [module k V] : Module k := Module.of k V +-- (4) is trickier, +-- and we would probably want to formalise as an equivalence of categories, +-- because "it's hard to get back to where we started". +example (X : Type*) : Module (free_algebra k X) ≃ Σ (V : Module k), X → (V ⟶ V) := skipped + +-- Definition 2.3.4 +-- A subrepresentation can be described using `submodule`, +variables (S : submodule A M) +-- or using the category theory library either as a monomorphism +variables (S' : Module.{u} A) (i : S' ⟶ N) [mono i] +-- or a subobject (defined as an isomorphism class of monomorphisms) +variables (S'' : subobject N) + +-- Definition 2.3.5: We express that a representation is "irreducible" using `simple`. +example (N : Module A) : Prop := simple N +-- there's also a predicate for unbundled modules: +example : simple (Module.of A M) ↔ is_simple_module A M := simple_iff_is_simple_module + +-- Definition 2.3.6: homomorphisms, intertwiners, isomorphisms +-- For unbundled representations, we use linear maps: +variables {M' : Type*} [add_comm_group M'] [module k M'] [module A M'] [is_scalar_tower k A M'] +variables (f : M →ₗ[A] M') +-- while for bundled representations we use the categorical morphism arrow: +variables (N₁ N₂ : Module.{u} A) (g : N₁ ⟶ N₂) +-- For isomorphisms, use one of +variables (e : M ≃ₗ[A] M') (j : N₁ ≅ N₂) + +-- Definition 2.3.7: direct sum +example : module A (M × M') := by apply_instance +example (N₁ N₂ : Module.{u} A) : Module.{u} A := N₁ ⊞ N₂ +example (N₁ N₂ : Module.{u} A) : N₁ ⊞ N₂ ≅ Module.of A (N₁ × N₂) := Module.biprod_iso_prod N₁ N₂ + +-- Definition 2.3.8: indecomposable +example (N : Module A) : Prop := indecomposable N +example (N : Module A) [simple N] : indecomposable N := indecomposable_of_simple N + +-- Proposition 2.3.9 (Schur's lemma) +example (N₁ N₂ : Module.{u} A) [simple N₁] (f : N₁ ⟶ N₂) (w : f ≠ 0) : mono f := +mono_of_nonzero_from_simple w +example (N₁ N₂ : Module.{u} A) [simple N₂] (f : N₁ ⟶ N₂) (w : f ≠ 0) : epi f := +epi_of_nonzero_to_simple w +example (N₁ N₂ : Module.{u} A) [simple N₁] [simple N₂] (f : N₁ ⟶ N₂) (w : f ≠ 0) : is_iso f := +is_iso_of_hom_simple w + +-- Corollary 2.3.10 (Schur's lemma over an algebraically closed field) +-- Unfortunately these can't be global instances +example [is_alg_closed k] (V : Module.{u} A) [simple V] [finite_dimensional k V] (f : V ⟶ V) : + ∃ φ : k, φ • 𝟙 V = f := +endomorphism_simple_eq_smul_id k f +-- Note that some magic is going on behind the scenes in this proof. +-- We're using a version of Schur's lemma that applies to any `k`-linear category, +-- and its hypotheses include `finite_dimensional k (V ⟶ V)` +-- rather than `finite_dimensional k V` (because `V` need not even be a vector space). +-- Typeclass inference is automatically generating this fact. + +-- Remark 2.3.11 (Schur's lemma doesn't hold over a non-algebraically closed field) +example : simple (Module.of ℂ ℂ) := simple_of_finrank_eq_one (finrank_self ℂ) +example : finite_dimensional ℝ (Module.of ℂ ℂ) := by { dsimp, apply_instance, } +example : + let V := Module.of ℂ ℂ in + ∃ (f : V ⟶ V), ∀ φ : ℝ, (φ : ℂ) • 𝟙 V ≠ f := +⟨algebra.lsmul ℂ ℂ complex.I, + λ φ w, by simpa using congr_arg complex.im (linear_map.congr_fun w (1 : ℂ))⟩ + +-- Corollary 2.3.12 +-- Every irreducible finite dimensional representation of a commutative algebra is 1-dimensional +example (A : Type*) [comm_ring A] [algebra k A] (V : Module A) [finite_dimensional k V] [simple V] : + finrank k V = 1 := +skipped + +-- Remark 2.3.13: Every 1-dimensional representation is irreducible +example (V : Module A) [finite_dimensional k V] (h : finrank k V = 1) : simple V := +simple_of_finrank_eq_one h + +-- Example 2.3.14: skipped (1 and 3 we can do, 2 requires Jordan normal form) + +/-! +## 2.4 "Ideals" +-/ + +-- To be continued... diff --git a/src/algebra/category/Module/algebra.lean b/src/algebra/category/Module/algebra.lean index 1851f61081c1d..224baa250b672 100644 --- a/src/algebra/category/Module/algebra.lean +++ b/src/algebra/category/Module/algebra.lean @@ -25,7 +25,7 @@ these instances will not necessarily agree with the original ones. It seems without making a parallel version `Module' k A`, for modules over a `k`-algebra `A`, that carries these typeclasses, this seems hard to achieve. -(An alternative would be to always require these typeclasses, +(An alternative would be to always require these typeclasses, and remove the original `Module`, requiring users to write `Module' ℤ A` when `A` is merely a ring.) -/ diff --git a/src/linear_algebra/matrix/to_lin.lean b/src/linear_algebra/matrix/to_lin.lean index 06025cd606198..af53c60451b29 100644 --- a/src/linear_algebra/matrix/to_lin.lean +++ b/src/linear_algebra/matrix/to_lin.lean @@ -742,7 +742,7 @@ 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. -/ +target are, as 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 _) From 229f6f14a8b345d28ad17aaa1e9e79beb9e231da Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Thu, 30 Mar 2023 00:32:06 +0000 Subject: [PATCH 0726/1291] feat(set_theory/zfc/basic): define Set and Class intersection (#18232) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Yaël Dillies Co-authored-by: Scott Morrison --- src/set_theory/zfc/basic.lean | 66 +++++++++++++++++++++++++++++++++-- 1 file changed, 63 insertions(+), 3 deletions(-) diff --git a/src/set_theory/zfc/basic.lean b/src/set_theory/zfc/basic.lean index 8a27059edaabc..cb6a6be94c5e7 100644 --- a/src/set_theory/zfc/basic.lean +++ b/src/set_theory/zfc/basic.lean @@ -669,21 +669,52 @@ resp.eval 1 ⟨pSet.sUnion, λ ⟨α, A⟩ ⟨β, B⟩ ⟨αβ, βα⟩, prefix (name := Set.sUnion) `⋃₀ `:110 := Set.sUnion +/-- The intersection operator, the collection of elements in all of the elements of a ZFC set. We +special-case `⋂₀ ∅ = ∅`. -/ +noncomputable def sInter (x : Set) : Set := +by { classical, exact dite x.nonempty (λ h, {y ∈ h.some | ∀ z ∈ x, y ∈ z}) (λ _, ∅) } + +prefix (name := Set.sInter) `⋂₀ `:110 := Set.sInter + @[simp] theorem mem_sUnion {x y : Set.{u}} : y ∈ ⋃₀ x ↔ ∃ z ∈ x, y ∈ z := quotient.induction_on₂ x y (λ x y, iff.trans mem_sUnion ⟨λ ⟨z, h⟩, ⟨⟦z⟧, h⟩, λ ⟨z, h⟩, quotient.induction_on z (λ z h, ⟨z, h⟩) h⟩) +theorem mem_sInter {x y : Set} (h : x.nonempty) : y ∈ ⋂₀ x ↔ ∀ z ∈ x, y ∈ z := +begin + rw [sInter, dif_pos h], + simp only [mem_to_set, mem_sep, and_iff_right_iff_imp], + exact λ H, H _ h.some_mem +end + +@[simp] theorem sUnion_empty : ⋃₀ (∅ : Set) = ∅ := by { ext, simp } +@[simp] theorem sInter_empty : ⋂₀ (∅ : Set) = ∅ := dif_neg $ by simp + +theorem mem_of_mem_sInter {x y z : Set} (hy : y ∈ ⋂₀ x) (hz : z ∈ x) : y ∈ z := +begin + rcases eq_empty_or_nonempty x with rfl | hx, + { exact (not_mem_empty z hz).elim }, + { exact (mem_sInter hx).1 hy z hz } +end + theorem mem_sUnion_of_mem {x y z : Set} (hy : y ∈ z) (hz : z ∈ x) : y ∈ ⋃₀ x := mem_sUnion.2 ⟨z, hz, hy⟩ -@[simp] theorem sUnion_empty : ⋃₀ (∅ : Set.{u}) = ∅ := by { ext, simp } +theorem not_mem_sInter_of_not_mem {x y z : Set} (hy : ¬ y ∈ z) (hz : z ∈ x) : ¬ y ∈ ⋂₀ x := +λ hx, hy $ mem_of_mem_sInter hx hz @[simp] theorem sUnion_singleton {x : Set.{u}} : ⋃₀ ({x} : Set) = x := ext $ λ y, by simp_rw [mem_sUnion, exists_prop, mem_singleton, exists_eq_left] +@[simp] theorem sInter_singleton {x : Set.{u}} : ⋂₀ ({x} : Set) = x := +ext $ λ y, by simp_rw [mem_sInter (singleton_nonempty x), mem_singleton, forall_eq] + @[simp] theorem to_set_sUnion (x : Set.{u}) : (⋃₀ x).to_set = ⋃₀ (to_set '' x.to_set) := by { ext, simp } +theorem to_set_sInter {x : Set.{u}} (h : x.nonempty) : (⋂₀ x).to_set = ⋂₀ (to_set '' x.to_set) := +by { ext, simp [mem_sInter h] } + theorem singleton_injective : function.injective (@singleton Set Set _) := λ x y H, let this := congr_arg sUnion H in by rwa [sUnion_singleton, sUnion_singleton] at this @@ -1009,6 +1040,11 @@ def sUnion (x : Class) : Class := ⋃₀ (Class_to_Cong x) prefix (name := Class.sUnion) `⋃₀ `:110 := Class.sUnion +/-- The intersection of a class is the class of all members of ZFC sets in the class -/ +def sInter (x : Class) : Class := ⋂₀ Class_to_Cong x + +prefix (name := Class.sInter) `⋂₀ `:110 := Class.sInter + theorem of_Set.inj {x y : Set.{u}} (h : (x : Class.{u}) = y) : x = y := Set.ext $ λ z, by { change (x : Class.{u}) z ↔ (y : Class.{u}) z, rw h } @@ -1067,8 +1103,32 @@ begin exact ⟨z, rfl, w, hwx, hwz⟩ } end -@[simp] theorem sUnion_empty : ⋃₀ (∅ : Class.{u}) = (∅ : Class.{u}) := -by { ext, simp } +@[simp] theorem sInter_apply {x : Class.{u}} {y : Set.{u}} : + (⋂₀ x) y ↔ ∀ z : Set.{u}, x z → y ∈ z := +begin + refine ⟨λ hxy z hxz, hxy _ ⟨z, rfl, hxz⟩, _⟩, + rintro H - ⟨z, rfl, hxz⟩, + exact H _ hxz +end + +@[simp, norm_cast] theorem sInter_coe {x : Set.{u}} (h : x.nonempty) : ⋂₀ (x : Class.{u}) = ⋂₀ x := +set.ext $ λ y, sInter_apply.trans (Set.mem_sInter h).symm + +theorem mem_of_mem_sInter {x y z : Class} (hy : y ∈ ⋂₀ x) (hz : z ∈ x) : y ∈ z := +by { obtain ⟨w, rfl, hw⟩ := hy, exact coe_mem.2 (hw z hz) } + +theorem mem_sInter {x y : Class.{u}} (h : x.nonempty) : y ∈ ⋂₀ x ↔ ∀ z, z ∈ x → y ∈ z := +begin + refine ⟨λ hy z, mem_of_mem_sInter hy, λ H, _⟩, + simp_rw [mem_def, sInter_apply], + obtain ⟨z, hz⟩ := h, + obtain ⟨y, rfl, hzy⟩ := H z (coe_mem.2 hz), + refine ⟨y, rfl, λ w hxw, _⟩, + simpa only [coe_mem, coe_apply] using H w (coe_mem.2 hxw), +end + +@[simp] theorem sUnion_empty : ⋃₀ (∅ : Class.{u}) = ∅ := by { ext, simp } +@[simp] theorem sInter_empty : ⋂₀ (∅ : Class.{u}) = univ := by { ext, simp [sInter, ←univ] } /-- An induction principle for sets. If every subset of a class is a member, then the class is universal. -/ From 564bcc44d2b394a50c0cd6340c14a6b02a50a99a Mon Sep 17 00:00:00 2001 From: Reid Barton Date: Thu, 30 Mar 2023 00:32:08 +0000 Subject: [PATCH 0727/1291] fix(algebra.group.ulift): fix to_additive name for cancel_comm_monoid (#18268) This is a fix to match mathlib 4 https://leanprover-community.github.io/mathlib4_docs/Mathlib/Algebra/Group/ULift.html#ULift.addCancelCommMonoid --- src/algebra/group/ulift.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/algebra/group/ulift.lean b/src/algebra/group/ulift.lean index c90c9083e88f2..7635fe79241ee 100644 --- a/src/algebra/group/ulift.lean +++ b/src/algebra/group/ulift.lean @@ -161,7 +161,7 @@ instance cancel_monoid [cancel_monoid α] : cancel_monoid (ulift α) := equiv.ulift.injective.cancel_monoid _ rfl (λ _ _, rfl) (λ _ _, rfl) -@[to_additive add_cancel_monoid] +@[to_additive add_cancel_comm_monoid] instance cancel_comm_monoid [cancel_comm_monoid α] : cancel_comm_monoid (ulift α) := equiv.ulift.injective.cancel_comm_monoid _ rfl (λ _ _, rfl) (λ _ _, rfl) From c04bc6e93e23aa0182aba53661a2211e80b6feac Mon Sep 17 00:00:00 2001 From: Amelia Livingston <101damnations@github.com> Date: Thu, 30 Mar 2023 00:32:10 +0000 Subject: [PATCH 0728/1291] refactor(representation_theory/group_cohomology_resolution): refactor `k[G^{n + 1}]` isomorphism (#18271) This refactors the isomorphism $k[G^{n + 1}] \cong k[G] \otimes_k k[G^n]$ (where $G$ acts by left multiplication on $k[G^{n + 1}]$ and $k[G]$ but trivially on $k[G^n]$) to use an isomorphism of $G$-sets $G^{n + 1} \cong G \times G^n.$ --- src/representation_theory/Action.lean | 68 ++++- src/representation_theory/Rep.lean | 71 ++++- src/representation_theory/basic.lean | 8 +- .../group_cohomology_resolution.lean | 247 +++++++++--------- 4 files changed, 254 insertions(+), 140 deletions(-) diff --git a/src/representation_theory/Action.lean b/src/representation_theory/Action.lean index bb5e99b5489d4..21c36f2b712a8 100644 --- a/src/representation_theory/Action.lean +++ b/src/representation_theory/Action.lean @@ -14,6 +14,7 @@ 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.monoidal.types import category_theory.abelian.functor_category import category_theory.abelian.transfer import category_theory.conj @@ -405,6 +406,13 @@ begin simp, end +/-- Given an object `X` isomorphic to the tensor unit of `V`, `X` equipped with the trivial action +is isomorphic to the tensor unit of `Action V G`. -/ +def tensor_unit_iso {X : V} (f : 𝟙_ V ≅ X) : + 𝟙_ (Action V G) ≅ Action.mk X 1 := +Action.mk_iso f (λ g, by simp only [monoid_hom.one_apply, End.one_def, category.id_comp f.hom, + tensor_unit_rho, category.comp_id]) + variables (V G) /-- When `V` is monoidal the forgetful functor `Action V G` to `V` is monoidal. -/ @@ -653,6 +661,47 @@ def of_mul_action_limit_cone {ι : Type v} (G : Type (max v u)) [monoid G] congr, end } } +/-- The `G`-set `G`, acting on itself by left multiplication. -/ +@[simps] def left_regular (G : Type u) [monoid G] : Action (Type u) (Mon.of G) := +Action.of_mul_action G G + +/-- The `G`-set `Gⁿ`, acting on itself by left multiplication. -/ +@[simps] def diagonal (G : Type u) [monoid G] (n : ℕ) : Action (Type u) (Mon.of G) := +Action.of_mul_action G (fin n → G) + +/-- We have `fin 1 → G ≅ G` as `G`-sets, with `G` acting by left multiplication. -/ +def diagonal_one_iso_left_regular (G : Type u) [monoid G] : + diagonal G 1 ≅ left_regular G := Action.mk_iso (equiv.fun_unique _ _).to_iso (λ g, rfl) + +/-- Given `X : Action (Type u) (Mon.of G)` for `G` a group, then `G × X` (with `G` acting as left +multiplication on the first factor and by `X.ρ` on the second) is isomorphic as a `G`-set to +`G × X` (with `G` acting as left multiplication on the first factor and trivially on the second). +The isomorphism is given by `(g, x) ↦ (g, g⁻¹ • x)`. -/ +@[simps] def left_regular_tensor_iso (G : Type u) [group G] + (X : Action (Type u) (Mon.of G)) : + left_regular G ⊗ X ≅ left_regular G ⊗ Action.mk X.V 1 := +{ hom := + { hom := λ g, ⟨g.1, (X.ρ (g.1⁻¹ : G) g.2 : X.V)⟩, + comm' := λ g, funext $ λ x, prod.ext rfl $ + show (X.ρ ((g * x.1)⁻¹ : G) * X.ρ g) x.2 = _, + by simpa only [mul_inv_rev, ←X.ρ.map_mul, inv_mul_cancel_right] }, + inv := + { hom := λ g, ⟨g.1, X.ρ g.1 g.2⟩, + comm' := λ g, funext $ λ x, prod.ext rfl $ + by simpa only [tensor_rho, types_comp_apply, tensor_apply, left_regular_ρ_apply, map_mul] }, + hom_inv_id' := hom.ext _ _ (funext $ λ x, prod.ext rfl $ + show (X.ρ x.1 * X.ρ (x.1⁻¹ : G)) x.2 = _, + by simpa only [←X.ρ.map_mul, mul_inv_self, X.ρ.map_one]), + inv_hom_id' := hom.ext _ _ (funext $ λ x, prod.ext rfl $ + show (X.ρ (x.1⁻¹ : G) * X.ρ x.1) _ = _, + by simpa only [←X.ρ.map_mul, inv_mul_self, X.ρ.map_one]) } + +/-- The natural isomorphism of `G`-sets `Gⁿ⁺¹ ≅ G × Gⁿ`, where `G` acts by left multiplication on +each factor. -/ +@[simps] def diagonal_succ (G : Type u) [monoid G] (n : ℕ) : + diagonal G (n + 1) ≅ left_regular G ⊗ diagonal G n := +mk_iso (equiv.pi_fin_succ_above_equiv _ 0).to_iso (λ g, rfl) + end Action namespace category_theory.functor @@ -689,10 +738,11 @@ namespace category_theory.monoidal_functor open Action variables {V} {W : Type (u+1)} [large_category W] [monoidal_category V] [monoidal_category W] + (F : monoidal_functor V W) (G : Mon.{u}) /-- A monoidal functor induces a monoidal functor between the categories of `G`-actions within those categories. -/ -@[simps] def map_Action (F : monoidal_functor V W) (G : Mon.{u}) : +@[simps] def map_Action : monoidal_functor (Action V G) (Action W G) := { ε := { hom := F.ε, @@ -709,4 +759,20 @@ the categories of `G`-actions within those categories. -/ right_unitality' := by { intros, ext, dsimp, simp, dsimp, simp, }, ..F.to_functor.map_Action G, } +@[simp] lemma map_Action_ε_inv_hom : + (inv (F.map_Action G).ε).hom = inv F.ε := +begin + ext, + simp only [←F.map_Action_to_lax_monoidal_functor_ε_hom G, ←Action.comp_hom, + is_iso.hom_inv_id, id_hom], +end + +@[simp] lemma map_Action_μ_inv_hom (X Y : Action V G) : + (inv ((F.map_Action G).μ X Y)).hom = inv (F.μ X.V Y.V) := +begin + ext, + simpa only [←F.map_Action_to_lax_monoidal_functor_μ_hom G, ←Action.comp_hom, + is_iso.hom_inv_id, id_hom], +end + end category_theory.monoidal_functor diff --git a/src/representation_theory/Rep.lean b/src/representation_theory/Rep.lean index 20544f6b3366d..b0cf5eea56fbe 100644 --- a/src/representation_theory/Rep.lean +++ b/src/representation_theory/Rep.lean @@ -76,6 +76,17 @@ lemma of_ρ_apply {V : Type u} [add_comm_group V] [module k V] (ρ : representation k G V) (g : Mon.of G) : (Rep.of ρ).ρ g = ρ (g : G) := rfl +variables (k G) + +/-- The trivial `k`-linear `G`-representation on a `k`-module `V.` -/ +def trivial (V : Type u) [add_comm_group V] [module k V] : Rep k G := +Rep.of (@representation.trivial k G V _ _ _ _) + +variables {k G} + +lemma trivial_def {V : Type u} [add_comm_group V] [module k V] (g : G) (v : V) : + (trivial k G V).ρ g v = v := rfl + -- Verify that limits are calculated correctly. noncomputable example : preserves_limits (forget₂ (Rep k G) (Module.{u} k)) := by apply_instance @@ -101,23 +112,55 @@ noncomputable def linearization : variables {k G} @[simp] lemma linearization_obj_ρ (X : Action (Type u) (Mon.of G)) (g : G) (x : X.V →₀ k) : - ((linearization k G).1.1.obj X).ρ g x = finsupp.lmap_domain k k (X.ρ g) x := rfl + ((linearization k G).obj X).ρ g x = finsupp.lmap_domain k k (X.ρ g) x := rfl @[simp] lemma linearization_of (X : Action (Type u) (Mon.of G)) (g : G) (x : X.V) : - ((linearization k G).1.1.obj X).ρ g (finsupp.single x (1 : k)) + ((linearization k G).obj X).ρ g (finsupp.single x (1 : k)) = finsupp.single (X.ρ g x) (1 : k) := by rw [linearization_obj_ρ, finsupp.lmap_domain_apply, finsupp.map_domain_single] -variables (X Y : Action (Type u) (Mon.of G)) (f : X ⟶ Y) +variables {X Y : Action (Type u) (Mon.of G)} (f : X ⟶ Y) @[simp] lemma linearization_map_hom : - ((linearization k G).1.1.map f).hom = finsupp.lmap_domain k k f.hom := rfl + ((linearization k G).map f).hom = finsupp.lmap_domain k k f.hom := rfl -lemma linearization_map_hom_of (x : X.V) : - ((linearization k G).1.1.map f).hom (finsupp.single x (1 : k)) - = finsupp.single (f.hom x) (1 : k) := +lemma linearization_map_hom_single (x : X.V) (r : k) : + ((linearization k G).map f).hom (finsupp.single x r) + = finsupp.single (f.hom x) r := by rw [linearization_map_hom, finsupp.lmap_domain_apply, finsupp.map_domain_single] +@[simp] lemma linearization_μ_hom (X Y : Action (Type u) (Mon.of G)) : + ((linearization k G).μ X Y).hom = (finsupp_tensor_finsupp' k X.V Y.V).to_linear_map := +rfl + +@[simp] lemma linearization_μ_inv_hom (X Y : Action (Type u) (Mon.of G)) : + (inv ((linearization k G).μ X Y)).hom = (finsupp_tensor_finsupp' k X.V Y.V).symm.to_linear_map := +begin + simp_rw [←Action.forget_map, functor.map_inv, Action.forget_map, linearization_μ_hom], + apply is_iso.inv_eq_of_hom_inv_id _, + exact linear_map.ext (λ x, linear_equiv.symm_apply_apply _ _), +end + +@[simp] lemma linearization_ε_hom : + (linearization k G).ε.hom = finsupp.lsingle punit.star := +rfl + +@[simp] lemma linearization_ε_inv_hom_apply (r : k) : + (inv (linearization k G).ε).hom (finsupp.single punit.star r) = r := +begin + simp_rw [←Action.forget_map, functor.map_inv, Action.forget_map], + rw [←finsupp.lsingle_apply punit.star r], + apply is_iso.hom_inv_id_apply _ _, +end + +variables (k G) + +/-- The linearization of a type `X` on which `G` acts trivially is the trivial `G`-representation +on `k[X]`. -/ +@[simps] noncomputable def linearization_trivial_iso (X : Type u) : + (linearization k G).obj (Action.mk X 1) ≅ trivial k G (X →₀ k) := +Action.mk_iso (iso.refl _) $ λ g, by { ext1, ext1, exact linearization_of _ _ _ } + variables (k G) /-- Given a `G`-action on `H`, this is `k[H]` bundled with the natural representation @@ -125,11 +168,19 @@ variables (k G) noncomputable abbreviation of_mul_action (H : Type u) [mul_action G H] : Rep k G := of $ representation.of_mul_action k G H +/-- The `k`-linear `G`-representation on `k[G]`, induced by left multiplication. -/ +noncomputable def left_regular : Rep k G := +of_mul_action k G G + +/-- The `k`-linear `G`-representation on `k[Gⁿ]`, induced by left multiplication. -/ +noncomputable def diagonal (n : ℕ) : Rep k G := +of_mul_action k G (fin n → G) + /-- The linearization of a type `H` with a `G`-action is definitionally isomorphic to the `k`-linear `G`-representation on `k[H]` induced by the `G`-action on `H`. -/ -noncomputable def linearization_of_mul_action_iso (n : ℕ) : - (linearization k G).1.1.obj (Action.of_mul_action G (fin n → G)) - ≅ of_mul_action k G (fin n → G) := iso.refl _ +noncomputable def linearization_of_mul_action_iso (H : Type u) [mul_action G H] : + (linearization k G).obj (Action.of_mul_action G H) + ≅ of_mul_action k G H := iso.refl _ variables {k G} diff --git a/src/representation_theory/basic.lean b/src/representation_theory/basic.lean index 81292da3bc65b..9bef318df98be 100644 --- a/src/representation_theory/basic.lean +++ b/src/representation_theory/basic.lean @@ -47,15 +47,15 @@ namespace representation section trivial -variables {k G V : Type*} [comm_semiring k] [monoid G] [add_comm_monoid V] [module k V] +variables (k : Type*) {G V : Type*} [comm_semiring k] [monoid G] [add_comm_monoid V] [module k V] /-- -The trivial representation of `G` on the one-dimensional module `k`. +The trivial representation of `G` on a `k`-module V. -/ -def trivial : representation k G k := 1 +def trivial : representation k G V := 1 @[simp] -lemma trivial_def (g : G) (v : k) : trivial g v = v := rfl +lemma trivial_def (g : G) (v : V) : trivial k g v = v := rfl end trivial diff --git a/src/representation_theory/group_cohomology_resolution.lean b/src/representation_theory/group_cohomology_resolution.lean index 987be7a8fe044..aae7db290557f 100644 --- a/src/representation_theory/group_cohomology_resolution.lean +++ b/src/representation_theory/group_cohomology_resolution.lean @@ -37,10 +37,8 @@ standard projective resolution of `k` as a trivial `k`-linear `G`-representation ## Main definitions - * `group_cohomology.resolution.to_tensor` - * `group_cohomology.resolution.of_tensor` - * `Rep.of_mul_action` - * `group_cohomology.resolution.equiv_tensor` + * `group_cohomology.resolution.Action_diagonal_succ` + * `group_cohomology.resolution.diagonal_succ` * `group_cohomology.resolution.of_mul_action_basis` * `classifying_space_universal_cover` * `group_cohomology.resolution.forget₂_to_Module_homotopy_equiv` @@ -60,12 +58,10 @@ over `k`. -/ noncomputable theory - universes u v w variables {k G : Type u} [comm_ring k] {n : ℕ} -open_locale tensor_product open category_theory local notation `Gⁿ` := fin n → G @@ -73,132 +69,134 @@ local notation `Gⁿ⁺¹` := fin (n + 1) → G namespace group_cohomology.resolution -open finsupp (hiding lift) fin (partial_prod) representation +open finsupp (hiding lift) fin (partial_prod) section basis variables (k G n) [group G] -/-- The `k`-linear map from `k[Gⁿ⁺¹]` to `k[G] ⊗ₖ k[Gⁿ]` sending `(g₀, ..., gₙ)` -to `g₀ ⊗ (g₀⁻¹g₁, g₁⁻¹g₂, ..., gₙ₋₁⁻¹gₙ)`. -/ -def to_tensor_aux : - ((fin (n + 1) → G) →₀ k) →ₗ[k] (G →₀ k) ⊗[k] ((fin n → G) →₀ k) := -finsupp.lift ((G →₀ k) ⊗[k] ((fin n → G) →₀ k)) k (fin (n + 1) → G) - (λ x, single (x 0) 1 ⊗ₜ[k] single (λ i, (x i)⁻¹ * x i.succ) 1) - -/-- The `k`-linear map from `k[G] ⊗ₖ k[Gⁿ]` to `k[Gⁿ⁺¹]` sending `g ⊗ (g₁, ..., gₙ)` to -`(g, gg₁, gg₁g₂, ..., gg₁...gₙ)`. -/ -def of_tensor_aux : - (G →₀ k) ⊗[k] ((fin n → G) →₀ k) →ₗ[k] ((fin (n + 1) → G) →₀ k) := -tensor_product.lift (finsupp.lift _ _ _ $ λ g, finsupp.lift _ _ _ - (λ f, single (g • partial_prod f) (1 : k))) - -variables {k G n} +section Action +open Action + +/-- An isomorphism of `G`-sets `Gⁿ⁺¹ ≅ G × Gⁿ`, where `G` acts by left multiplication on `Gⁿ⁺¹` and +`G` but trivially on `Gⁿ`. The map sends `(g₀, ..., gₙ) ↦ (g₀, (g₀⁻¹g₁, g₁⁻¹g₂, ..., gₙ₋₁⁻¹gₙ))`, +and the inverse is `(g₀, (g₁, ..., gₙ)) ↦ (g₀, g₀g₁, g₀g₁g₂, ..., g₀g₁...gₙ).` -/ +def Action_diagonal_succ (G : Type u) [group G] : Π (n : ℕ), + diagonal G (n + 1) ≅ left_regular G ⊗ Action.mk (fin n → G) 1 +| 0 := diagonal_one_iso_left_regular G ≪≫ (ρ_ _).symm ≪≫ tensor_iso (iso.refl _) + (tensor_unit_iso (equiv.equiv_of_unique punit _).to_iso) +| (n+1) := diagonal_succ _ _ ≪≫ tensor_iso (iso.refl _) (Action_diagonal_succ n) + ≪≫ left_regular_tensor_iso _ _ ≪≫ tensor_iso (iso.refl _) (mk_iso + (equiv.pi_fin_succ_above_equiv (λ j, G) 0).symm.to_iso (λ g, rfl)) + +lemma Action_diagonal_succ_hom_apply {G : Type u} [group G] {n : ℕ} (f : fin (n + 1) → G) : + (Action_diagonal_succ G n).hom.hom f = (f 0, λ i, (f i)⁻¹ * f i.succ) := +begin + induction n with n hn, + { exact prod.ext rfl (funext $ λ x, fin.elim0 x) }, + { ext, + { refl }, + { dunfold Action_diagonal_succ, + simp only [iso.trans_hom, comp_hom, types_comp_apply, diagonal_succ_hom_hom, + left_regular_tensor_iso_hom_hom, tensor_iso_hom, mk_iso_hom_hom, equiv.to_iso_hom, + tensor_hom, equiv.pi_fin_succ_above_equiv_symm_apply, tensor_apply, types_id_apply, + tensor_rho, monoid_hom.one_apply, End.one_def, hn (λ (j : fin (n + 1)), f j.succ), + fin.coe_eq_cast_succ, fin.insert_nth_zero'], + refine fin.cases (fin.cons_zero _ _) (λ i, _) x, + { simp only [fin.cons_succ, mul_left_inj, inv_inj, fin.cast_succ_fin_succ], }}} +end -lemma to_tensor_aux_single (f : Gⁿ⁺¹) (m : k) : - to_tensor_aux k G n (single f m) = single (f 0) m ⊗ₜ single (λ i, (f i)⁻¹ * f i.succ) 1 := +lemma Action_diagonal_succ_inv_apply {G : Type u} [group G] {n : ℕ} (g : G) (f : fin n → G) : + (Action_diagonal_succ G n).inv.hom (g, f) = (g • fin.partial_prod f : fin (n + 1) → G) := begin - simp only [to_tensor_aux, lift_apply, sum_single_index, tensor_product.smul_tmul'], - { simp }, + revert g, + induction n with n hn, + { intros g, + ext, + simpa only [subsingleton.elim x 0, pi.smul_apply, fin.partial_prod_zero, + smul_eq_mul, mul_one] }, + { intro g, + ext, + dunfold Action_diagonal_succ, + simp only [iso.trans_inv, comp_hom, hn, diagonal_succ_inv_hom, types_comp_apply, + tensor_iso_inv, iso.refl_inv, tensor_hom, id_hom, tensor_apply, types_id_apply, + left_regular_tensor_iso_inv_hom, tensor_rho, left_regular_ρ_apply, pi.smul_apply, + smul_eq_mul], + refine fin.cases _ _ x, + { simp only [fin.cons_zero, fin.partial_prod_zero, mul_one], }, + { intro i, + simpa only [fin.cons_succ, pi.smul_apply, smul_eq_mul, fin.partial_prod_succ', mul_assoc], }} end -lemma to_tensor_aux_of_mul_action (g : G) (x : Gⁿ⁺¹) : - to_tensor_aux k G n (of_mul_action k G Gⁿ⁺¹ g (single x 1)) = - tensor_product.map (of_mul_action k G G g) 1 (to_tensor_aux k G n (single x 1)) := -by simp [of_mul_action_def, to_tensor_aux_single, mul_assoc, inv_mul_cancel_left] +end Action +section Rep +open Rep + +/-- An isomorphism of `k`-linear representations of `G` from `k[Gⁿ⁺¹]` to `k[G] ⊗ₖ k[Gⁿ]` (on +which `G` acts by `ρ(g₁)(g₂ ⊗ x) = (g₁ * g₂) ⊗ x`) sending `(g₀, ..., gₙ)` to +`g₀ ⊗ (g₀⁻¹g₁, g₁⁻¹g₂, ..., gₙ₋₁⁻¹gₙ)`. The inverse sends `g₀ ⊗ (g₁, ..., gₙ)` to +`(g₀, g₀g₁, ..., g₀g₁...gₙ)`. -/ +def diagonal_succ (n : ℕ) : + diagonal k G (n + 1) ≅ left_regular k G ⊗ trivial k G ((fin n → G) →₀ k) := +(linearization k G).map_iso (Action_diagonal_succ G n) + ≪≫ (as_iso ((linearization k G).μ (Action.left_regular G) _)).symm + ≪≫ tensor_iso (iso.refl _) (linearization_trivial_iso k G (fin n → G)) -lemma of_tensor_aux_single (g : G) (m : k) (x : Gⁿ →₀ k) : - of_tensor_aux k G n ((single g m) ⊗ₜ x) = - finsupp.lift (Gⁿ⁺¹ →₀ k) k Gⁿ (λ f, single (g • partial_prod f) m) x := -by simp [of_tensor_aux, sum_single_index, smul_sum, mul_comm m] +variables {k G n} -lemma of_tensor_aux_comm_of_mul_action (g h : G) (x : Gⁿ) : - of_tensor_aux k G n (tensor_product.map (of_mul_action k G G g) - (1 : module.End k (Gⁿ →₀ k)) (single h (1 : k) ⊗ₜ single x (1 : k))) = - of_mul_action k G Gⁿ⁺¹ g (of_tensor_aux k G n (single h 1 ⊗ₜ single x 1)) := +lemma diagonal_succ_hom_single (f : Gⁿ⁺¹) (a : k) : + (diagonal_succ k G n).hom.hom (single f a) = + single (f 0) 1 ⊗ₜ single (λ i, (f i)⁻¹ * f i.succ) a := begin - simp [of_mul_action_def, of_tensor_aux_single, mul_smul], + dunfold diagonal_succ, + simpa only [iso.trans_hom, iso.symm_hom, Action.comp_hom, Module.comp_def, linear_map.comp_apply, + functor.map_iso_hom, linearization_map_hom_single (Action_diagonal_succ G n).hom f a, + as_iso_inv, linearization_μ_inv_hom, Action_diagonal_succ_hom_apply, finsupp_tensor_finsupp', + linear_equiv.trans_symm, lcongr_symm, linear_equiv.trans_apply, lcongr_single, + tensor_product.lid_symm_apply, finsupp_tensor_finsupp_symm_single, + linear_equiv.coe_to_linear_map], end -lemma to_tensor_aux_left_inv (x : Gⁿ⁺¹ →₀ k) : - of_tensor_aux _ _ _ (to_tensor_aux _ _ _ x) = x := +lemma diagonal_succ_inv_single_single (g : G) (f : Gⁿ) (a b : k) : + (diagonal_succ k G n).inv.hom (finsupp.single g a ⊗ₜ finsupp.single f b) = + single (g • partial_prod f) (a * b) := begin - refine linear_map.ext_iff.1 (@finsupp.lhom_ext _ _ _ k _ _ _ _ _ - (linear_map.comp (of_tensor_aux _ _ _) (to_tensor_aux _ _ _)) linear_map.id (λ x y, _)) x, - dsimp, - rw [to_tensor_aux_single x y, of_tensor_aux_single, finsupp.lift_apply, finsupp.sum_single_index, - one_smul, fin.partial_prod_left_inv], - { rw zero_smul } + dunfold diagonal_succ, + simp only [iso.trans_inv, iso.symm_inv, iso.refl_inv, tensor_iso_inv, Action.tensor_hom, + Action.comp_hom, Module.comp_def, linear_map.comp_apply, as_iso_hom, functor.map_iso_inv, + Module.monoidal_category.hom_apply, linearization_trivial_iso_inv_hom_apply, + linearization_μ_hom, Action.id_hom ((linearization k G).obj _), Action_diagonal_succ_inv_apply, + Module.id_apply, linear_equiv.coe_to_linear_map, + finsupp_tensor_finsupp'_single_tmul_single k (Action.left_regular G).V, + linearization_map_hom_single (Action_diagonal_succ G n).inv (g, f) (a * b)], end -lemma to_tensor_aux_right_inv (x : (G →₀ k) ⊗[k] (Gⁿ →₀ k)) : - to_tensor_aux _ _ _ (of_tensor_aux _ _ _ x) = x := +lemma diagonal_succ_inv_single_left (g : G) (f : Gⁿ →₀ k) (r : k) : + (diagonal_succ k G n).inv.hom (finsupp.single g r ⊗ₜ f) = + finsupp.lift (Gⁿ⁺¹ →₀ k) k Gⁿ (λ f, single (g • partial_prod f) r) f := begin - refine tensor_product.induction_on x (by simp) (λ y z, _) (λ z w hz hw, by simp [hz, hw]), - rw [←finsupp.sum_single y, finsupp.sum, tensor_product.sum_tmul], - simp only [finset.smul_sum, linear_map.map_sum], - refine finset.sum_congr rfl (λ f hf, _), - simp only [of_tensor_aux_single, finsupp.lift_apply, finsupp.smul_single', - linear_map.map_finsupp_sum, to_tensor_aux_single, fin.partial_prod_right_inv], - dsimp, - simp only [fin.partial_prod_zero, mul_one], - conv_rhs {rw [←finsupp.sum_single z, finsupp.sum, tensor_product.tmul_sum]}, - exact finset.sum_congr rfl (λ g hg, show _ ⊗ₜ _ = _, by - rw [←finsupp.smul_single', tensor_product.smul_tmul, finsupp.smul_single_one]) + refine f.induction _ _, + { simp only [tensor_product.tmul_zero, map_zero] }, + { intros a b x ha hb hx, + simp only [lift_apply, smul_single', mul_one, tensor_product.tmul_add, map_add, + diagonal_succ_inv_single_single, hx, finsupp.sum_single_index, + mul_comm b, zero_mul, single_zero] }, end -variables (k G n) - -/-- A hom of `k`-linear representations of `G` from `k[Gⁿ⁺¹]` to `k[G] ⊗ₖ k[Gⁿ]` (on which `G` acts -by `ρ(g₁)(g₂ ⊗ x) = (g₁ * g₂) ⊗ x`) sending `(g₀, ..., gₙ)` to -`g₀ ⊗ (g₀⁻¹g₁, g₁⁻¹g₂, ..., gₙ₋₁⁻¹gₙ)`. -/ -def to_tensor : Rep.of_mul_action k G (fin (n + 1) → G) ⟶ Rep.of - ((representation.of_mul_action k G G).tprod (1 : G →* module.End k ((fin n → G) →₀ k))) := -{ hom := to_tensor_aux k G n, - comm' := λ g, by ext; exact to_tensor_aux_of_mul_action _ _ } - -/-- A hom of `k`-linear representations of `G` from `k[G] ⊗ₖ k[Gⁿ]` (on which `G` acts -by `ρ(g₁)(g₂ ⊗ x) = (g₁ * g₂) ⊗ x`) to `k[Gⁿ⁺¹]` sending `g ⊗ (g₁, ..., gₙ)` to -`(g, gg₁, gg₁g₂, ..., gg₁...gₙ)`. -/ -def of_tensor : - Rep.of ((representation.of_mul_action k G G).tprod (1 : G →* module.End k ((fin n → G) →₀ k))) - ⟶ Rep.of_mul_action k G (fin (n + 1) → G) := -{ hom := of_tensor_aux k G n, - comm' := λ g, by { ext, congr' 1, exact (of_tensor_aux_comm_of_mul_action _ _ _) }} - -variables {k G n} - -@[simp] lemma to_tensor_single (f : Gⁿ⁺¹) (m : k) : - (to_tensor k G n).hom (single f m) = single (f 0) m ⊗ₜ single (λ i, (f i)⁻¹ * f i.succ) 1 := -by apply to_tensor_aux_single f m - -@[simp] lemma of_tensor_single (g : G) (m : k) (x : Gⁿ →₀ k) : - (of_tensor k G n).hom ((single g m) ⊗ₜ x) = - finsupp.lift (Rep.of_mul_action k G Gⁿ⁺¹) k Gⁿ (λ f, single (g • partial_prod f) m) x := -by apply of_tensor_aux_single g m x - -lemma of_tensor_single' (g : G →₀ k) (x : Gⁿ) (m : k) : - (of_tensor k G n).hom (g ⊗ₜ single x m) = - finsupp.lift _ k G (λ a, single (a • partial_prod x) m) g := -by simp [of_tensor, of_tensor_aux] +lemma diagonal_succ_inv_single_right (g : G →₀ k) (f : Gⁿ) (r : k) : + (diagonal_succ k G n).inv.hom (g ⊗ₜ finsupp.single f r) = + finsupp.lift _ k G (λ a, single (a • partial_prod f) r) g := +begin + refine g.induction _ _, + { simp only [tensor_product.zero_tmul, map_zero], }, + { intros a b x ha hb hx, + simp only [lift_apply, smul_single', map_add, hx, diagonal_succ_inv_single_single, + tensor_product.add_tmul, finsupp.sum_single_index, zero_mul, single_zero] } +end +end Rep variables (k G n) - -/-- An isomorphism of `k`-linear representations of `G` from `k[Gⁿ⁺¹]` to `k[G] ⊗ₖ k[Gⁿ]` (on -which `G` acts by `ρ(g₁)(g₂ ⊗ x) = (g₁ * g₂) ⊗ x`) sending `(g₀, ..., gₙ)` to -`g₀ ⊗ (g₀⁻¹g₁, g₁⁻¹g₂, ..., gₙ₋₁⁻¹gₙ)`. -/ -def equiv_tensor : (Rep.of_mul_action k G (fin (n + 1) → G)) ≅ Rep.of - ((representation.of_mul_action k G G).tprod (1 : representation k G ((fin n → G) →₀ k))) := -Action.mk_iso (linear_equiv.to_Module_iso -{ inv_fun := of_tensor_aux k G n, - left_inv := to_tensor_aux_left_inv, - right_inv := λ x, by convert to_tensor_aux_right_inv x, - ..to_tensor_aux k G n }) (to_tensor k G n).comm - -@[simp] lemma equiv_tensor_def : - (equiv_tensor k G n).hom = to_tensor k G n := rfl - -@[simp] lemma equiv_tensor_inv_def : - (equiv_tensor k G n).inv = of_tensor k G n := rfl +open_locale tensor_product +open representation /-- The `k[G]`-linear isomorphism `k[G] ⊗ₖ k[Gⁿ] ≃ k[Gⁿ⁺¹]`, where the `k[G]`-module structure on the lefthand side is `tensor_product.left_module`, whilst that of the righthand side comes from @@ -217,7 +215,7 @@ def of_mul_action_basis_aux : (monoid_algebra k G ⊗[k] ((fin n → G) →₀ k rw [←of_mul_action_self_smul_eq_mul, smul_tprod_one_as_module] }, { rw [smul_add, hz, hy, smul_add], } end, .. ((Rep.equivalence_Module_monoid_algebra.1).map_iso - (equiv_tensor k G n).symm).to_linear_equiv } + (diagonal_succ k G n).symm).to_linear_equiv } /-- A `k[G]`-basis of `k[Gⁿ⁺¹]`, coming from the `k[G]`-linear isomorphism `k[G] ⊗ₖ k[Gⁿ] ≃ k[Gⁿ⁺¹].` -/ @@ -241,7 +239,7 @@ open group_cohomology.resolution `Hom(k[Gⁿ⁺¹], A)` is `k`-linearly isomorphic to the set of functions `Gⁿ → A`. -/ noncomputable def diagonal_hom_equiv (A : Rep k G) : (Rep.of_mul_action k G (fin (n + 1) → G) ⟶ A) ≃ₗ[k] ((fin n → G) → A) := -linear.hom_congr k ((equiv_tensor k G n).trans +linear.hom_congr k ((diagonal_succ k G n).trans ((representation.of_mul_action k G G).Rep_of_tprod_iso 1)) (iso.refl _) ≪≫ₗ ((Rep.monoidal_closed.linear_hom_equiv_comm _ _ _) ≪≫ₗ (Rep.left_regular_hom_equiv _)) ≪≫ₗ (finsupp.llift A k k (fin n → G)).symm @@ -259,9 +257,8 @@ begin simpa only [linear_equiv.trans_apply, Rep.left_regular_hom_equiv_apply, monoidal_closed.linear_hom_equiv_comm_hom, finsupp.llift_symm_apply, tensor_product.curry_apply, linear.hom_congr_apply, iso.refl_hom, iso.trans_inv, Action.comp_hom, Module.comp_def, - linear_map.comp_apply, equiv_tensor_inv_def, representation.Rep_of_tprod_iso_inv_apply, - of_tensor_single (1 : G) (1 : k) (finsupp.single x (1 : k)), finsupp.lift_apply, - finsupp.sum_single_index, one_smul, zero_smul], + linear_map.comp_apply, representation.Rep_of_tprod_iso_inv_apply, + diagonal_succ_inv_single_single (1 : G) x, one_smul, one_mul], end /-- Given a `k`-linear `G`-representation `A`, `diagonal_hom_equiv` is a `k`-linear isomorphism of @@ -278,10 +275,10 @@ begin Rep.left_regular_hom_equiv_symm_apply, linear.hom_congr_symm_apply, Action.comp_hom, iso.refl_inv, category.comp_id, Rep.monoidal_closed.linear_hom_equiv_comm_symm_hom, iso.trans_hom, Module.comp_def, linear_map.comp_apply, representation.Rep_of_tprod_iso_apply, - equiv_tensor_def, to_tensor_single x (1 : k), tensor_product.uncurry_apply, - Rep.left_regular_hom_hom, finsupp.lift_apply, Rep.ihom_obj_ρ, representation.lin_hom_apply, - finsupp.sum_single_index, zero_smul, one_smul, Rep.of_ρ, - monoid_hom.one_apply, linear_map.one_apply, finsupp.llift_apply A k k], + diagonal_succ_hom_single x (1 : k), tensor_product.uncurry_apply, Rep.left_regular_hom_hom, + finsupp.lift_apply, Rep.ihom_obj_ρ, representation.lin_hom_apply, finsupp.sum_single_index, + zero_smul, one_smul, Rep.of_ρ, Rep.Action_ρ_eq_ρ, Rep.trivial_def (x 0)⁻¹, + finsupp.llift_apply A k k], end end Rep @@ -434,7 +431,7 @@ homotopy equivalent to the complex which is `k` at 0 and 0 elsewhere. -/ def forget₂_to_Module_homotopy_equiv : homotopy_equiv (group_cohomology.resolution.forget₂_to_Module k G) ((chain_complex.single₀ (Module k)).obj - ((forget₂ (Rep k G) _).obj $ Rep.of representation.trivial)) := + ((forget₂ (Rep k G) _).obj $ Rep.trivial k G k)) := (homotopy_equiv.of_iso (comp_forget_augmented_iso k G).symm).trans $ (simplicial_object.augmented.extra_degeneracy.homotopy_equiv (extra_degeneracy_comp_forget_augmented_to_Module k G)).trans @@ -443,7 +440,7 @@ def forget₂_to_Module_homotopy_equiv : homotopy_equiv types.terminal_iso.to_equiv.unique).to_Module_iso) /-- The hom of `k`-linear `G`-representations `k[G¹] → k` sending `∑ nᵢgᵢ ↦ ∑ nᵢ`. -/ -def ε : Rep.of_mul_action k G (fin 1 → G) ⟶ Rep.of representation.trivial := +def ε : Rep.of_mul_action k G (fin 1 → G) ⟶ Rep.trivial k G k := { hom := finsupp.total _ _ _ (λ f, (1 : k)), comm' := λ g, begin @@ -493,7 +490,7 @@ end /-- The chain map from the standard resolution of `k` to `k[0]` given by `∑ nᵢgᵢ ↦ ∑ nᵢ` in degree zero. -/ def ε_to_single₀ : group_cohomology.resolution k G ⟶ (chain_complex.single₀ _).obj - (Rep.of representation.trivial) := + (Rep.trivial k G k) := ((group_cohomology.resolution k G).to_single₀_equiv _).symm ⟨ε k G, d_comp_ε k G⟩ lemma ε_to_single₀_comp_eq : ((forget₂ _ (Module.{u} k)).map_homological_complex _).map @@ -526,7 +523,7 @@ variables [group G] /-- The standard projective resolution of `k` as a trivial `k`-linear `G`-representation. -/ def group_cohomology.ProjectiveResolution : - ProjectiveResolution (Rep.of (@representation.trivial k G _ _)) := + ProjectiveResolution (Rep.trivial k G k) := (ε_to_single₀ k G).to_single₀_ProjectiveResolution (X_projective k G) instance : enough_projectives (Rep k G) := @@ -537,7 +534,7 @@ Rep.equivalence_Module_monoid_algebra.enough_projectives_iff.2 `G`-representation) is isomorphic to the `n`th cohomology group of `Hom(P, V)`, where `P` is the standard resolution of `k` called `group_cohomology.resolution k G`. -/ def group_cohomology.Ext_iso (V : Rep k G) (n : ℕ) : - ((Ext k (Rep k G) n).obj (opposite.op $ Rep.of representation.trivial)).obj V ≅ + ((Ext k (Rep k G) n).obj (opposite.op $ Rep.trivial k G k)).obj V ≅ (((((linear_yoneda k (Rep k G)).obj V).right_op.map_homological_complex _).obj (group_cohomology.resolution k G)).homology n).unop := by let := (((linear_yoneda k (Rep k G)).obj V).right_op.left_derived_obj_iso From 91862a6001a8b6ae3f261cdd8eea42f6ac596886 Mon Sep 17 00:00:00 2001 From: sgouezel Date: Thu, 30 Mar 2023 00:32:12 +0000 Subject: [PATCH 0729/1291] feat(analysis/calculus/cont_diff): bound for the iterated derivative of a product (#18632) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We show that `‖D^n (x ↦ B (f x) (g x))‖ ≤ ‖B‖ ∑_{k ≤ n} n.choose k ‖D^k f‖ ‖D^{n-k} g‖` for a general continuous bilinear map `B`, and specialize this inequality to the case of the product of two functions. This is a first step to control the norm of the iterated derivative of a composition. --- src/analysis/calculus/cont_diff.lean | 340 ++++++++++++++++++- src/analysis/normed_space/operator_norm.lean | 3 + 2 files changed, 342 insertions(+), 1 deletion(-) diff --git a/src/analysis/calculus/cont_diff.lean b/src/analysis/calculus/cont_diff.lean index a0b45dce2eb49..5ef0d2d58d2a5 100644 --- a/src/analysis/calculus/cont_diff.lean +++ b/src/analysis/calculus/cont_diff.lean @@ -30,15 +30,73 @@ open_locale classical big_operators nnreal local notation `∞` := (⊤ : ℕ∞) -universes u v w uE uF uG +universes u v w uD uE uF uG local attribute [instance, priority 1001] normed_add_comm_group.to_add_comm_group normed_space.to_module' add_comm_group.to_add_comm_monoid +namespace finset + +/- TODO porting note: move the next two lemmas to the file `data.nat.choose.sum` -/ +/-- The sum of `(n+1).choose i * f i (n+1-i)` can be split into two sums at rank `n`, +respectively of `n.choose i * f i (n+1-i)` and `n.choose i * f (i+1) (n-i)`. -/ +lemma sum_choose_succ_mul {R : Type*} [semiring R] (f : ℕ → ℕ → R) (n : ℕ) : + ∑ i in range (n+2), ((n+1).choose i : R) * f i (n + 1 - i) = + ∑ i in range (n+1), (n.choose i : R) * f i (n + 1 - i) + + ∑ i in range (n+1), (n.choose i : R) * f (i + 1) (n - i) := +begin + have A : ∑ i in range (n + 1), (n.choose (i+1) : R) * f (i + 1) (n - i) + f 0 (n + 1) + = ∑ i in range (n+1), n.choose i * f i (n + 1 - i), + { rw [finset.sum_range_succ, finset.sum_range_succ'], + simp only [nat.choose_succ_self, algebra_map.coe_zero, zero_mul, add_zero, + nat.succ_sub_succ_eq_sub, nat.choose_zero_right, algebra_map.coe_one, one_mul, tsub_zero] }, + calc + ∑ i in finset.range (n+2), ((n+1).choose i : R) * f i (n + 1 - i) + = ∑ i in finset.range (n+1), ((n+1).choose (i+1) : R) * f (i+1) (n + 1 - (i+1)) + + f 0 (n + 1 - 0) : + begin + rw finset.sum_range_succ', + simp only [nat.choose_zero_right, algebra_map.coe_one, one_mul], + end + ... = ∑ i in finset.range (n+1), (n.choose i : R) * f i (n + 1 - i) + + ∑ i in finset.range (n+1), n.choose i * f (i + 1) (n - i) : + begin + simp only [nat.choose_succ_succ, nat.cast_add, nat.succ_sub_succ_eq_sub, tsub_zero, add_mul], + rw [finset.sum_add_distrib, ← A], + abel, + end +end + +/-- The sum along the antidiagonal of `(n+1).choose i * f i j` can be split into two sums along the +antidiagonal at rank `n`, respectively of `n.choose i * f i (j+1)` and `n.choose j * f (i+1) j`. -/ +lemma sum_antidiagonal_choose_succ_mul {R : Type*} [semiring R] (f : ℕ → ℕ → R) (n : ℕ) : + ∑ ij in nat.antidiagonal (n + 1), ((n + 1).choose ij.1 : R) * f ij.1 ij.2 = + ∑ ij in nat.antidiagonal n, (n.choose ij.1 : R) * f ij.1 (ij.2 + 1) + + ∑ ij in nat.antidiagonal n, (n.choose ij.2 : R) * f (ij.1 + 1) ij.2 := +begin + convert sum_choose_succ_mul f n using 1, + { exact nat.sum_antidiagonal_eq_sum_range_succ (λ i j, ((n+1).choose i : R) * f i j) (n+1) }, + congr' 1, + { rw nat.sum_antidiagonal_eq_sum_range_succ (λ i j, (n.choose i : R) * f i (j + 1)) n, + apply finset.sum_congr rfl (λ i hi, _), + have : n + 1 - i = n - i + 1, from nat.sub_add_comm (nat.lt_succ_iff.1 (finset.mem_range.1 hi)), + simp only [this] }, + { suffices H : ∑ ij in nat.antidiagonal n, (n.choose ij.2 : R) * f (ij.1 + 1) ij.2 + = ∑ ij in nat.antidiagonal n, (n.choose ij.1 : R) * f (ij.1 + 1) ij.2, + by rw [H, nat.sum_antidiagonal_eq_sum_range_succ (λ i j, (n.choose i : R) * f (i + 1) j) n], + apply finset.sum_congr rfl (λ i hi, _), + congr' 2, + apply nat.choose_symm_of_eq_add, + rw [← nat.mem_antidiagonal.1 hi, add_comm] } +end + +end finset + open set fin filter function open_locale topology variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] +{D : Type uD} [normed_add_comm_group D] [normed_space 𝕜 D] {E : Type uE} [normed_add_comm_group E] [normed_space 𝕜 E] {F : Type uF} [normed_add_comm_group F] [normed_space 𝕜 F] {G : Type uG} [normed_add_comm_group G] [normed_space 𝕜 G] @@ -2264,3 +2322,283 @@ lemma cont_diff.restrict_scalars (h : cont_diff 𝕜' n f) : cont_diff_iff_cont_diff_at.2 $ λ x, h.cont_diff_at.restrict_scalars _ end restrict_scalars + +/-!## Quantitative bounds -/ + +/-- Bounding the norm of the iterated derivative of `B (f x) (g x)` within a set in terms of the +iterated derivatives of `f` and `g` when `B` is bilinear. This lemma is an auxiliary version +assuming all spaces live in the same universe, to enable an induction. Use instead +`continuous_linear_map.norm_iterated_fderiv_within_le_of_bilinear` that removes this assumption. -/ +lemma continuous_linear_map.norm_iterated_fderiv_within_le_of_bilinear_aux + {Du Eu Fu Gu : Type u} + [normed_add_comm_group Du] [normed_space 𝕜 Du] + [normed_add_comm_group Eu] [normed_space 𝕜 Eu] + [normed_add_comm_group Fu] [normed_space 𝕜 Fu] + [normed_add_comm_group Gu] [normed_space 𝕜 Gu] + (B : Eu →L[𝕜] Fu →L[𝕜] Gu) {f : Du → Eu} {g : Du → Fu} {n : ℕ} {s : set Du} {x : Du} + (hf : cont_diff_on 𝕜 n f s) (hg : cont_diff_on 𝕜 n g s) (hs : unique_diff_on 𝕜 s) (hx : x ∈ s) : + ‖iterated_fderiv_within 𝕜 n (λ y, B (f y) (g y)) s x‖ + ≤ ‖B‖ * ∑ i in finset.range (n+1), (n.choose i : ℝ) + * ‖iterated_fderiv_within 𝕜 i f s x‖ * ‖iterated_fderiv_within 𝕜 (n-i) g s x‖ := +begin + /- We argue by induction on `n`. The bound is trivial for `n = 0`. For `n + 1`, we write + the `(n+1)`-th derivative as the `n`-th derivative of the derivative `B f g' + B f' g`, and apply + the inductive assumption to each of those two terms. For this induction to make sense, + the spaces of linear maps that appear in the induction should be in the same universe as the + original spaces, which explains why we assume in the lemma that all spaces live in the same + universe. -/ + unfreezingI { induction n with n IH generalizing Eu Fu Gu}, + { simp only [←mul_assoc, norm_iterated_fderiv_within_zero, finset.range_one, finset.sum_singleton, + nat.choose_self, algebra_map.coe_one, one_mul], + apply ((B (f x)).le_op_norm (g x)).trans, + apply mul_le_mul_of_nonneg_right _ (norm_nonneg _), + exact B.le_op_norm (f x) }, + { have In : (n : with_top ℕ) + 1 ≤ n.succ, by simp only [nat.cast_succ, le_refl], + have I1 : + ‖iterated_fderiv_within 𝕜 n (λ (y : Du), B.precompR Du (f y) (fderiv_within 𝕜 g s y)) s x‖ ≤ + ‖B‖ * ∑ (i : ℕ) in finset.range (n + 1), n.choose i * + ‖iterated_fderiv_within 𝕜 i f s x‖ * ‖iterated_fderiv_within 𝕜 (n + 1 - i) g s x‖ := calc + ‖iterated_fderiv_within 𝕜 n (λ (y : Du), B.precompR Du (f y) (fderiv_within 𝕜 g s y)) s x‖ + ≤ ‖B.precompR Du‖ * ∑ (i : ℕ) in finset.range (n + 1), n.choose i + * ‖iterated_fderiv_within 𝕜 i f s x‖ + * ‖iterated_fderiv_within 𝕜 (n - i) (fderiv_within 𝕜 g s) s x‖ : + IH _ (hf.of_le (nat.cast_le.2 (nat.le_succ n))) (hg.fderiv_within hs In) + ... ≤ ‖B‖ * ∑ (i : ℕ) in finset.range (n + 1), n.choose i + * ‖iterated_fderiv_within 𝕜 i f s x‖ + * ‖iterated_fderiv_within 𝕜 (n - i) (fderiv_within 𝕜 g s) s x‖ : + mul_le_mul_of_nonneg_right (B.norm_precompR_le Du) (finset.sum_nonneg' (λ i, by positivity)) + ... = _ : + begin + congr' 1, + apply finset.sum_congr rfl (λ i hi, _ ), + rw [nat.succ_sub (nat.lt_succ_iff.1 (finset.mem_range.1 hi)), + iterated_fderiv_within_succ_eq_comp_right hs hx, linear_isometry_equiv.norm_map], + end, + have I2 : + ‖iterated_fderiv_within 𝕜 n (λ (y : Du), B.precompL Du (fderiv_within 𝕜 f s y) (g y)) s x‖ ≤ + ‖B‖ * ∑ (i : ℕ) in finset.range (n + 1), n.choose i * + ‖iterated_fderiv_within 𝕜 (i + 1) f s x‖ * ‖iterated_fderiv_within 𝕜 (n - i) g s x‖ := calc + ‖iterated_fderiv_within 𝕜 n (λ (y : Du), B.precompL Du (fderiv_within 𝕜 f s y) (g y)) s x‖ + ≤ ‖B.precompL Du‖ * ∑ (i : ℕ) in finset.range (n + 1), n.choose i + * ‖iterated_fderiv_within 𝕜 i (fderiv_within 𝕜 f s) s x‖ + * ‖iterated_fderiv_within 𝕜 (n - i) g s x‖ : + IH _ (hf.fderiv_within hs In) (hg.of_le (nat.cast_le.2 (nat.le_succ n))) + ... ≤ ‖B‖ * ∑ (i : ℕ) in finset.range (n + 1), n.choose i + * ‖iterated_fderiv_within 𝕜 i (fderiv_within 𝕜 f s) s x‖ + * ‖iterated_fderiv_within 𝕜 (n - i) g s x‖ : + mul_le_mul_of_nonneg_right (B.norm_precompL_le Du) (finset.sum_nonneg' (λ i, by positivity)) + ... = _ : + begin + congr' 1, + apply finset.sum_congr rfl (λ i hi, _ ), + rw [iterated_fderiv_within_succ_eq_comp_right hs hx, linear_isometry_equiv.norm_map], + end, + have J : iterated_fderiv_within 𝕜 n + (λ (y : Du), fderiv_within 𝕜 (λ (y : Du), B (f y) (g y)) s y) s x + = iterated_fderiv_within 𝕜 n (λ y, B.precompR Du (f y) (fderiv_within 𝕜 g s y) + + B.precompL Du (fderiv_within 𝕜 f s y) (g y)) s x, + { apply iterated_fderiv_within_congr hs (λ y hy, _) hx, + have L : (1 : with_top ℕ) ≤ n.succ, + by simpa only [enat.coe_one, nat.one_le_cast] using nat.succ_pos n, + exact B.fderiv_within_of_bilinear (hf.differentiable_on L y hy) + (hg.differentiable_on L y hy) (hs y hy) }, + rw [iterated_fderiv_within_succ_eq_comp_right hs hx, linear_isometry_equiv.norm_map, J], + have A : cont_diff_on 𝕜 n (λ y, B.precompR Du (f y) (fderiv_within 𝕜 g s y)) s, + from (B.precompR Du).is_bounded_bilinear_map.cont_diff.comp_cont_diff_on₂ + (hf.of_le (nat.cast_le.2 (nat.le_succ n))) (hg.fderiv_within hs In), + have A' : cont_diff_on 𝕜 n (λ y, B.precompL Du (fderiv_within 𝕜 f s y) (g y)) s, + from (B.precompL Du).is_bounded_bilinear_map.cont_diff.comp_cont_diff_on₂ + (hf.fderiv_within hs In) (hg.of_le (nat.cast_le.2 (nat.le_succ n))), + rw iterated_fderiv_within_add_apply' A A' hs hx, + apply (norm_add_le _ _).trans ((add_le_add I1 I2).trans (le_of_eq _)), + simp_rw [← mul_add, mul_assoc], + congr' 1, + exact (finset.sum_choose_succ_mul (λ i j, ‖iterated_fderiv_within 𝕜 i f s x‖ * + ‖iterated_fderiv_within 𝕜 j g s x‖) n).symm } +end + +/-- Bounding the norm of the iterated derivative of `B (f x) (g x)` within a set in terms of the +iterated derivatives of `f` and `g` when `B` is bilinear: +`‖D^n (x ↦ B (f x) (g x))‖ ≤ ‖B‖ ∑_{k ≤ n} n.choose k ‖D^k f‖ ‖D^{n-k} g‖` -/ +lemma continuous_linear_map.norm_iterated_fderiv_within_le_of_bilinear + (B : E →L[𝕜] F →L[𝕜] G) {f : D → E} {g : D → F} {N : with_top ℕ} {s : set D} {x : D} + (hf : cont_diff_on 𝕜 N f s) (hg : cont_diff_on 𝕜 N g s) (hs : unique_diff_on 𝕜 s) (hx : x ∈ s) + {n : ℕ} (hn : (n : with_top ℕ) ≤ N) : + ‖iterated_fderiv_within 𝕜 n (λ y, B (f y) (g y)) s x‖ + ≤ ‖B‖ * ∑ i in finset.range (n+1), (n.choose i : ℝ) + * ‖iterated_fderiv_within 𝕜 i f s x‖ * ‖iterated_fderiv_within 𝕜 (n-i) g s x‖ := +begin + /- We reduce the bound to the case where all spaces live in the same universe (in which we + already have proved the result), by using linear isometries between the spaces and their `ulift` + to a common universe. These linear isometries preserve the norm of the iterated derivative. -/ + let Du : Type (max uD uE uF uG) := ulift.{(max uE uF uG) uD} D, + let Eu : Type (max uD uE uF uG) := ulift.{(max uD uF uG) uE} E, + let Fu : Type (max uD uE uF uG) := ulift.{(max uD uE uG) uF} F, + let Gu : Type (max uD uE uF uG) := ulift.{(max uD uE uF) uG} G, + have isoD : Du ≃ₗᵢ[𝕜] D := linear_isometry_equiv.ulift 𝕜 D, + have isoE : Eu ≃ₗᵢ[𝕜] E := linear_isometry_equiv.ulift 𝕜 E, + have isoF : Fu ≃ₗᵢ[𝕜] F := linear_isometry_equiv.ulift 𝕜 F, + have isoG : Gu ≃ₗᵢ[𝕜] G := linear_isometry_equiv.ulift 𝕜 G, + -- lift `f` and `g` to versions `fu` and `gu` on the lifted spaces. + let fu : Du → Eu := isoE.symm ∘ f ∘ isoD, + let gu : Du → Fu := isoF.symm ∘ g ∘ isoD, + -- lift the bilinear map `B` to a bilinear map `Bu` on the lifted spaces. + let Bu₀ : Eu →L[𝕜] Fu →L[𝕜] G, + from ((B.comp (isoE : Eu →L[𝕜] E)).flip.comp (isoF : Fu →L[𝕜] F)).flip, + let Bu : Eu →L[𝕜] Fu →L[𝕜] Gu, from continuous_linear_map.compL 𝕜 Eu (Fu →L[𝕜] G) (Fu →L[𝕜] Gu) + (continuous_linear_map.compL 𝕜 Fu G Gu (isoG.symm : G →L[𝕜] Gu)) Bu₀, + have Bu_eq : (λ y, Bu (fu y) (gu y)) = isoG.symm ∘ (λ y, B (f y) (g y)) ∘ isoD, + { ext1 y, + simp only [Bu, continuous_linear_map.compL_apply, function.comp_app, + continuous_linear_map.coe_comp', linear_isometry_equiv.coe_coe'', + continuous_linear_map.flip_apply, linear_isometry_equiv.apply_symm_apply] }, + -- All norms are preserved by the lifting process. + have Bu_le : ‖Bu‖ ≤ ‖B‖, + { refine continuous_linear_map.op_norm_le_bound _ (norm_nonneg _) (λ y, _), + refine continuous_linear_map.op_norm_le_bound _ (by positivity) (λ x, _ ), + simp only [Bu, continuous_linear_map.compL_apply, continuous_linear_map.coe_comp', + function.comp_app, linear_isometry_equiv.coe_coe'', continuous_linear_map.flip_apply, + linear_isometry_equiv.norm_map], + calc ‖B (isoE y) (isoF x)‖ + ≤ ‖B (isoE y)‖ * ‖isoF x‖ : continuous_linear_map.le_op_norm _ _ + ... ≤ ‖B‖ * ‖isoE y‖ * ‖isoF x‖ : + mul_le_mul_of_nonneg_right (continuous_linear_map.le_op_norm _ _) (norm_nonneg _) + ... = ‖B‖ * ‖y‖ * ‖x‖ : by simp only [linear_isometry_equiv.norm_map] }, + let su := isoD ⁻¹' s, + have hsu : unique_diff_on 𝕜 su, + from isoD.to_continuous_linear_equiv.unique_diff_on_preimage_iff.2 hs, + let xu := isoD.symm x, + have hxu : xu ∈ su, + by simpa only [set.mem_preimage, linear_isometry_equiv.apply_symm_apply] using hx, + have xu_x : isoD xu = x, by simp only [linear_isometry_equiv.apply_symm_apply], + have hfu : cont_diff_on 𝕜 n fu su, from isoE.symm.cont_diff.comp_cont_diff_on + ((hf.of_le hn).comp_continuous_linear_map (isoD : Du →L[𝕜] D)), + have hgu : cont_diff_on 𝕜 n gu su, from isoF.symm.cont_diff.comp_cont_diff_on + ((hg.of_le hn).comp_continuous_linear_map (isoD : Du →L[𝕜] D)), + have Nfu : ∀ i, ‖iterated_fderiv_within 𝕜 i fu su xu‖ = ‖iterated_fderiv_within 𝕜 i f s x‖, + { assume i, + rw linear_isometry_equiv.norm_iterated_fderiv_within_comp_left _ _ hsu hxu, + rw [linear_isometry_equiv.norm_iterated_fderiv_within_comp_right _ _ hs, xu_x], + rwa ← xu_x at hx }, + have Ngu : ∀ i, ‖iterated_fderiv_within 𝕜 i gu su xu‖ = ‖iterated_fderiv_within 𝕜 i g s x‖, + { assume i, + rw linear_isometry_equiv.norm_iterated_fderiv_within_comp_left _ _ hsu hxu, + rw [linear_isometry_equiv.norm_iterated_fderiv_within_comp_right _ _ hs, xu_x], + rwa ← xu_x at hx }, + have NBu : ‖iterated_fderiv_within 𝕜 n (λ y, Bu (fu y) (gu y)) su xu‖ = + ‖iterated_fderiv_within 𝕜 n (λ y, B (f y) (g y)) s x‖, + { rw Bu_eq, + rw linear_isometry_equiv.norm_iterated_fderiv_within_comp_left _ _ hsu hxu, + rw [linear_isometry_equiv.norm_iterated_fderiv_within_comp_right _ _ hs, xu_x], + rwa ← xu_x at hx }, + -- state the bound for the lifted objects, and deduce the original bound from it. + have : ‖iterated_fderiv_within 𝕜 n (λ y, Bu (fu y) (gu y)) su xu‖ + ≤ ‖Bu‖ * ∑ i in finset.range (n+1), (n.choose i : ℝ) + * ‖iterated_fderiv_within 𝕜 i fu su xu‖ * ‖iterated_fderiv_within 𝕜 (n-i) gu su xu‖, + from Bu.norm_iterated_fderiv_within_le_of_bilinear_aux hfu hgu hsu hxu, + simp only [Nfu, Ngu, NBu] at this, + apply this.trans (mul_le_mul_of_nonneg_right Bu_le _), + exact finset.sum_nonneg' (λ i, by positivity), +end + +/-- Bounding the norm of the iterated derivative of `B (f x) (g x)` in terms of the +iterated derivatives of `f` and `g` when `B` is bilinear: +`‖D^n (x ↦ B (f x) (g x))‖ ≤ ‖B‖ ∑_{k ≤ n} n.choose k ‖D^k f‖ ‖D^{n-k} g‖` -/ +lemma continuous_linear_map.norm_iterated_fderiv_le_of_bilinear + (B : E →L[𝕜] F →L[𝕜] G) {f : D → E} {g : D → F} {N : with_top ℕ} + (hf : cont_diff 𝕜 N f) (hg : cont_diff 𝕜 N g) (x : D) + {n : ℕ} (hn : (n : with_top ℕ) ≤ N) : + ‖iterated_fderiv 𝕜 n (λ y, B (f y) (g y)) x‖ + ≤ ‖B‖ * ∑ i in finset.range (n+1), (n.choose i : ℝ) + * ‖iterated_fderiv 𝕜 i f x‖ * ‖iterated_fderiv 𝕜 (n-i) g x‖ := +begin + simp_rw [← iterated_fderiv_within_univ], + exact B.norm_iterated_fderiv_within_le_of_bilinear hf.cont_diff_on hg.cont_diff_on + unique_diff_on_univ (mem_univ x) hn, +end + +/-- Bounding the norm of the iterated derivative of `B (f x) (g x)` within a set in terms of the +iterated derivatives of `f` and `g` when `B` is bilinear of norm at most `1`: +`‖D^n (x ↦ B (f x) (g x))‖ ≤ ∑_{k ≤ n} n.choose k ‖D^k f‖ ‖D^{n-k} g‖` -/ +lemma continuous_linear_map.norm_iterated_fderiv_within_le_of_bilinear_of_le_one + (B : E →L[𝕜] F →L[𝕜] G) {f : D → E} {g : D → F} {N : with_top ℕ} {s : set D} {x : D} + (hf : cont_diff_on 𝕜 N f s) (hg : cont_diff_on 𝕜 N g s) (hs : unique_diff_on 𝕜 s) (hx : x ∈ s) + {n : ℕ} (hn : (n : with_top ℕ) ≤ N) (hB : ‖B‖ ≤ 1) : + ‖iterated_fderiv_within 𝕜 n (λ y, B (f y) (g y)) s x‖ + ≤ ∑ i in finset.range (n+1), (n.choose i : ℝ) + * ‖iterated_fderiv_within 𝕜 i f s x‖ * ‖iterated_fderiv_within 𝕜 (n-i) g s x‖ := +begin + apply (B.norm_iterated_fderiv_within_le_of_bilinear hf hg hs hx hn).trans, + apply mul_le_of_le_one_left (finset.sum_nonneg' (λ i, _)) hB, + positivity +end + +/-- Bounding the norm of the iterated derivative of `B (f x) (g x)` in terms of the +iterated derivatives of `f` and `g` when `B` is bilinear of norm at most `1`: +`‖D^n (x ↦ B (f x) (g x))‖ ≤ ∑_{k ≤ n} n.choose k ‖D^k f‖ ‖D^{n-k} g‖` -/ +lemma continuous_linear_map.norm_iterated_fderiv_le_of_bilinear_of_le_one + (B : E →L[𝕜] F →L[𝕜] G) {f : D → E} {g : D → F} {N : with_top ℕ} + (hf : cont_diff 𝕜 N f) (hg : cont_diff 𝕜 N g) (x : D) + {n : ℕ} (hn : (n : with_top ℕ) ≤ N) (hB : ‖B‖ ≤ 1) : + ‖iterated_fderiv 𝕜 n (λ y, B (f y) (g y)) x‖ + ≤ ∑ i in finset.range (n+1), (n.choose i : ℝ) + * ‖iterated_fderiv 𝕜 i f x‖ * ‖iterated_fderiv 𝕜 (n-i) g x‖ := +begin + simp_rw [← iterated_fderiv_within_univ], + exact B.norm_iterated_fderiv_within_le_of_bilinear_of_le_one hf.cont_diff_on hg.cont_diff_on + unique_diff_on_univ (mem_univ x) hn hB, +end + +section + +variables {𝕜' : Type*} [normed_field 𝕜'] [normed_algebra 𝕜 𝕜'] [normed_space 𝕜' F] + [is_scalar_tower 𝕜 𝕜' F] + +lemma norm_iterated_fderiv_within_smul_le + {f : E → 𝕜'} {g : E → F} {N : with_top ℕ} (hf : cont_diff_on 𝕜 N f s) (hg : cont_diff_on 𝕜 N g s) + (hs : unique_diff_on 𝕜 s) {x : E} (hx : x ∈ s) {n : ℕ} (hn : (n : with_top ℕ) ≤ N) : + ‖iterated_fderiv_within 𝕜 n (λ y, f y • g y) s x‖ + ≤ ∑ i in finset.range (n+1), (n.choose i : ℝ) + * ‖iterated_fderiv_within 𝕜 i f s x‖ * ‖iterated_fderiv_within 𝕜 (n-i) g s x‖ := +(continuous_linear_map.lsmul 𝕜 𝕜' : 𝕜' →L[𝕜] F →L[𝕜] F) + .norm_iterated_fderiv_within_le_of_bilinear_of_le_one hf hg hs hx hn + continuous_linear_map.op_norm_lsmul_le + +lemma norm_iterated_fderiv_smul_le + {f : E → 𝕜'} {g : E → F} {N : with_top ℕ} (hf : cont_diff 𝕜 N f) (hg : cont_diff 𝕜 N g) + (x : E) {n : ℕ} (hn : (n : with_top ℕ) ≤ N) : + ‖iterated_fderiv 𝕜 n (λ y, f y • g y) x‖ + ≤ ∑ i in finset.range (n+1), (n.choose i : ℝ) + * ‖iterated_fderiv 𝕜 i f x‖ * ‖iterated_fderiv 𝕜 (n-i) g x‖ := +(continuous_linear_map.lsmul 𝕜 𝕜' : 𝕜' →L[𝕜] F →L[𝕜] F) + .norm_iterated_fderiv_le_of_bilinear_of_le_one hf hg x hn + continuous_linear_map.op_norm_lsmul_le + +end + +section +variables {A : Type*} [normed_ring A] [normed_algebra 𝕜 A] + +lemma norm_iterated_fderiv_within_mul_le + {f : E → A} {g : E → A} {N : with_top ℕ} (hf : cont_diff_on 𝕜 N f s) (hg : cont_diff_on 𝕜 N g s) + (hs : unique_diff_on 𝕜 s) {x : E} (hx : x ∈ s) {n : ℕ} (hn : (n : with_top ℕ) ≤ N) : + ‖iterated_fderiv_within 𝕜 n (λ y, f y * g y) s x‖ + ≤ ∑ i in finset.range (n+1), (n.choose i : ℝ) + * ‖iterated_fderiv_within 𝕜 i f s x‖ * ‖iterated_fderiv_within 𝕜 (n-i) g s x‖ := +(continuous_linear_map.mul 𝕜 A : A →L[𝕜] A →L[𝕜] A) + .norm_iterated_fderiv_within_le_of_bilinear_of_le_one hf hg hs hx hn + (continuous_linear_map.op_norm_mul_le _ _) + +lemma norm_iterated_fderiv_mul_le + {f : E → A} {g : E → A} {N : with_top ℕ} (hf : cont_diff 𝕜 N f) (hg : cont_diff 𝕜 N g) + (x : E) {n : ℕ} (hn : (n : with_top ℕ) ≤ N) : + ‖iterated_fderiv 𝕜 n (λ y, f y * g y) x‖ + ≤ ∑ i in finset.range (n+1), (n.choose i : ℝ) + * ‖iterated_fderiv 𝕜 i f x‖ * ‖iterated_fderiv 𝕜 (n-i) g x‖ := +begin + simp_rw [← iterated_fderiv_within_univ], + exact norm_iterated_fderiv_within_mul_le hf.cont_diff_on + hg.cont_diff_on unique_diff_on_univ (mem_univ x) hn, +end + +end diff --git a/src/analysis/normed_space/operator_norm.lean b/src/analysis/normed_space/operator_norm.lean index 20b8d79b5e7d4..a2a0e71b7316e 100644 --- a/src/analysis/normed_space/operator_norm.lean +++ b/src/analysis/normed_space/operator_norm.lean @@ -922,6 +922,9 @@ def mul : 𝕜' →L[𝕜] 𝕜' →L[𝕜] 𝕜' := (linear_map.mul 𝕜 𝕜') @[simp] lemma op_norm_mul_apply_le (x : 𝕜') : ‖mul 𝕜 𝕜' x‖ ≤ ‖x‖ := (op_norm_le_bound _ (norm_nonneg x) (norm_mul_le x)) +lemma op_norm_mul_le : ‖mul 𝕜 𝕜'‖ ≤ 1 := +linear_map.mk_continuous₂_norm_le _ zero_le_one _ + /-- Simultaneous left- and right-multiplication in a non-unital normed algebra, considered as a continuous trilinear map. This is akin to its non-continuous version `linear_map.mul_left_right`, but there is a minor difference: `linear_map.mul_left_right` is uncurried. -/ From a6ece35404f60597c651689c1b46ead86de5ac1b Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 30 Mar 2023 00:32:13 +0000 Subject: [PATCH 0730/1291] chore(algebra/star/self_adjoint): generalize a lemma to `semifield` (#18687) This was missed in a previous commit, and backports a change made during forward-porting. --- src/algebra/star/self_adjoint.lean | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/src/algebra/star/self_adjoint.lean b/src/algebra/star/self_adjoint.lean index 8ebc3e5f2e3af..146c294221606 100644 --- a/src/algebra/star/self_adjoint.lean +++ b/src/algebra/star/self_adjoint.lean @@ -196,13 +196,13 @@ star_rat_cast _ end division_ring -section field -variables [field R] [star_ring R] +section semifield +variables [semifield R] [star_ring R] lemma div {x y : R} (hx : is_self_adjoint x) (hy : is_self_adjoint y) : is_self_adjoint (x / y) := by simp only [is_self_adjoint_iff, star_div', hx.star_eq, hy.star_eq] -end field +end semifield section has_smul variables [has_star R] [add_monoid A] [star_add_monoid A] [has_smul R A] [star_module R A] From 71b36b6f3bbe3b44e6538673819324d3ee9fcc96 Mon Sep 17 00:00:00 2001 From: Jeremy Tan Jie Rui Date: Thu, 30 Mar 2023 00:32:14 +0000 Subject: [PATCH 0731/1291] chore(order/category): Partially revert categories (#18690) It was pointed out in leanprover-community/mathlib4#3164 that certain changes in #18657 were unsound. This PR reverts those unsound changes. Co-authored-by: Parcly Taxel --- src/order/complete_boolean_algebra.lean | 2 +- src/order/hom/complete_lattice.lean | 6 +++--- 2 files changed, 4 insertions(+), 4 deletions(-) diff --git a/src/order/complete_boolean_algebra.lean b/src/order/complete_boolean_algebra.lean index 4108c483a0eeb..dd590f754bc04 100644 --- a/src/order/complete_boolean_algebra.lean +++ b/src/order/complete_boolean_algebra.lean @@ -18,7 +18,7 @@ distributive Boolean algebras. ## Typeclasses -* `order.frame`: Frm: A complete lattice whose `⊓` distributes over `⨆`. +* `order.frame`: Frame: A complete lattice whose `⊓` distributes over `⨆`. * `order.coframe`: Coframe: A complete lattice whose `⊔` distributes over `⨅`. * `complete_distrib_lattice`: Completely distributive lattices: A complete lattice whose `⊓` and `⊔` distribute over `⨆` and `⨅` respectively. diff --git a/src/order/hom/complete_lattice.lean b/src/order/hom/complete_lattice.lean index 548321502b4f4..354e1d03d6106 100644 --- a/src/order/hom/complete_lattice.lean +++ b/src/order/hom/complete_lattice.lean @@ -21,7 +21,7 @@ be satisfied by itself and all stricter types. * `Sup_hom`: Maps which preserve `⨆`. * `Inf_hom`: Maps which preserve `⨅`. -* `frame_hom`: Frm homomorphisms. Maps which preserve `⨆`, `⊓` and `⊤`. +* `frame_hom`: Frame homomorphisms. Maps which preserve `⨆`, `⊓` and `⊤`. * `complete_lattice_hom`: Complete lattice homomorphisms. Maps which preserve `⨆` and `⨅`. ## Typeclasses @@ -37,7 +37,7 @@ be satisfied by itself and all stricter types. ## TODO -Frm homs are Heyting homs. +Frame homs are Heyting homs. -/ open function order_dual set @@ -365,7 +365,7 @@ instance : order_top (Inf_hom α β) := ⟨⊤, λ f a, le_top⟩ end Inf_hom -/-! ### Frm homomorphisms -/ +/-! ### Frame homomorphisms -/ namespace frame_hom variables [complete_lattice α] [complete_lattice β] [complete_lattice γ] [complete_lattice δ] From 2ec920d35348cb2d13ac0e1a2ad9df0fdf1a76b4 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Thu, 30 Mar 2023 03:40:44 +0000 Subject: [PATCH 0732/1291] feat(data/list/lemmas): add lemmas about `set.range list.nth*` (#18647) Add versions for `list.nth_le`, `list.nth`, `list.nthd`, and `list.inth`. Also move lemmas from `list` to `set` namespace. --- src/data/list/lemmas.lean | 23 -------- src/data/multiset/basic.lean | 2 +- src/data/set/list.lean | 69 ++++++++++++++++++++++ src/group_theory/submonoid/membership.lean | 4 +- 4 files changed, 72 insertions(+), 26 deletions(-) create mode 100644 src/data/set/list.lean diff --git a/src/data/list/lemmas.lean b/src/data/list/lemmas.lean index 7705922f783c9..fc3df77ce584f 100644 --- a/src/data/list/lemmas.lean +++ b/src/data/list/lemmas.lean @@ -20,29 +20,6 @@ variables {α β γ : Type*} namespace list -lemma range_map (f : α → β) : set.range (map f) = {l | ∀ x ∈ l, x ∈ set.range f} := -begin - refine set.subset.antisymm (set.range_subset_iff.2 $ - λ l, forall_mem_map_iff.2 $ λ y _, set.mem_range_self _) (λ l hl, _), - induction l with a l ihl, { exact ⟨[], rfl⟩ }, - rcases ihl (λ x hx, hl x $ subset_cons _ _ hx) with ⟨l, rfl⟩, - rcases hl a (mem_cons_self _ _) with ⟨a, rfl⟩, - exact ⟨a :: l, map_cons _ _ _⟩ -end - -lemma range_map_coe (s : set α) : set.range (map (coe : s → α)) = {l | ∀ x ∈ l, x ∈ s} := -by rw [range_map, subtype.range_coe] - -/-- If each element of a list can be lifted to some type, then the whole list can be lifted to this -type. -/ -instance can_lift (c) (p) [can_lift α β c p] : - can_lift (list α) (list β) (list.map c) (λ l, ∀ x ∈ l, p x) := -{ prf := λ l H, - begin - rw [← set.mem_range, range_map], - exact λ a ha, can_lift.prf a (H a ha), - end} - lemma inj_on_insert_nth_index_of_not_mem (l : list α) (x : α) (hx : x ∉ l) : set.inj_on (λ k, insert_nth k x l) {n | n ≤ l.length} := begin diff --git a/src/data/multiset/basic.lean b/src/data/multiset/basic.lean index 72d9b9e8bf54c..ae6158c45359a 100644 --- a/src/data/multiset/basic.lean +++ b/src/data/multiset/basic.lean @@ -3,7 +3,7 @@ Copyright (c) 2015 Microsoft Corporation. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ -import data.list.lemmas +import data.set.list import data.list.perm /-! diff --git a/src/data/set/list.lean b/src/data/set/list.lean new file mode 100644 index 0000000000000..78ea83204dfd8 --- /dev/null +++ b/src/data/set/list.lean @@ -0,0 +1,69 @@ +/- +Copyright (c) 2023 Yury Kudryashov. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yury Kudryashov +-/ +import data.set.image +import data.list.basic +import data.fin.basic + +/-! +# Lemmas about `list`s and `set.range` + +In this file we prove lemmas about range of some operations on lists. +-/ + +open list +variables {α β : Type*} (l : list α) + +namespace set + +lemma range_list_map (f : α → β) : range (map f) = {l | ∀ x ∈ l, x ∈ range f} := +begin + refine subset.antisymm (range_subset_iff.2 $ λ l, forall_mem_map_iff.2 $ λ y _, mem_range_self _) + (λ l hl, _), + induction l with a l ihl, { exact ⟨[], rfl⟩ }, + rcases ihl (λ x hx, hl x $ subset_cons _ _ hx) with ⟨l, rfl⟩, + rcases hl a (mem_cons_self _ _) with ⟨a, rfl⟩, + exact ⟨a :: l, map_cons _ _ _⟩ +end + +lemma range_list_map_coe (s : set α) : range (map (coe : s → α)) = {l | ∀ x ∈ l, x ∈ s} := +by rw [range_list_map, subtype.range_coe] + +@[simp] lemma range_list_nth_le : range (λ k : fin l.length, l.nth_le k k.2) = {x | x ∈ l} := +begin + ext x, + rw [mem_set_of_eq, mem_iff_nth_le], + exact ⟨λ ⟨⟨n, h₁⟩, h₂⟩, ⟨n, h₁, h₂⟩, λ ⟨n, h₁, h₂⟩, ⟨⟨n, h₁⟩, h₂⟩⟩ +end + +lemma range_list_nth : range l.nth = insert none (some '' {x | x ∈ l}) := +begin + rw [← range_list_nth_le, ← range_comp], + refine (range_subset_iff.2 $ λ n, _).antisymm (insert_subset.2 ⟨_, _⟩), + exacts [(le_or_lt l.length n).imp nth_eq_none_iff.2 (λ hlt, ⟨⟨_, _⟩, (nth_le_nth hlt).symm⟩), + ⟨_, nth_eq_none_iff.2 le_rfl⟩, range_subset_iff.2 $ λ k, ⟨_, nth_le_nth _⟩] +end + +@[simp] lemma range_list_nthd (d : α) : range (λ n, l.nthd n d) = insert d {x | x ∈ l} := +calc range (λ n, l.nthd n d) = (λ o : option α, o.get_or_else d) '' range l.nth : + by simp only [← range_comp, (∘), nthd_eq_get_or_else_nth] +... = insert d {x | x ∈ l} : + by simp only [range_list_nth, image_insert_eq, option.get_or_else, image_image, image_id'] + +@[simp] +lemma range_list_inth [inhabited α] (l : list α) : range l.inth = insert default {x | x ∈ l} := +range_list_nthd l default + +end set + +/-- If each element of a list can be lifted to some type, then the whole list can be lifted to this +type. -/ +instance list.can_lift (c) (p) [can_lift α β c p] : + can_lift (list α) (list β) (list.map c) (λ l, ∀ x ∈ l, p x) := +{ prf := λ l H, + begin + rw [← set.mem_range, set.range_list_map], + exact λ a ha, can_lift.prf a (H a ha), + end} diff --git a/src/group_theory/submonoid/membership.lean b/src/group_theory/submonoid/membership.lean index b51c17fd96fc6..f275719bbacb0 100644 --- a/src/group_theory/submonoid/membership.lean +++ b/src/group_theory/submonoid/membership.lean @@ -294,8 +294,8 @@ by rw [free_monoid.mrange_lift, subtype.range_coe] @[to_additive] lemma closure_eq_image_prod (s : set M) : (closure s : set M) = list.prod '' {l : list M | ∀ x ∈ l, x ∈ s} := begin - rw [closure_eq_mrange, coe_mrange, ← list.range_map_coe, ← set.range_comp, function.comp], - exact congr_arg _ (funext $ free_monoid.lift_apply _), + rw [closure_eq_mrange, coe_mrange, ← set.range_list_map_coe, ← set.range_comp], + exact congr_arg _ (funext $ free_monoid.lift_apply _) end @[to_additive] From fe44cd36149e675eb5dec87acc7e8f1d6568e081 Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Thu, 30 Mar 2023 06:36:57 +0000 Subject: [PATCH 0733/1291] feat (number_theory/modular_forms): holomorphy of the Jacobi theta function (#18631) Show that the Jacobi theta function is holomorphic on the upper half-plane, with exponential decay at infinity. --- .../complex/locally_uniform_limit.lean | 35 ++++- .../complex/upper_half_plane/topology.lean | 18 ++- src/number_theory/modular_forms/basic.lean | 7 - .../modular_forms/jacobi_theta.lean | 138 ++++++++++++++---- 4 files changed, 164 insertions(+), 34 deletions(-) diff --git a/src/analysis/complex/locally_uniform_limit.lean b/src/analysis/complex/locally_uniform_limit.lean index 36e01d439ce67..26f8e16d780c0 100644 --- a/src/analysis/complex/locally_uniform_limit.lean +++ b/src/analysis/complex/locally_uniform_limit.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Vincent Beffara -/ import analysis.complex.removable_singularity -import analysis.calculus.uniform_limits_deriv +import analysis.calculus.series /-! # Locally uniform limits of holomorphic functions @@ -179,4 +179,37 @@ end end weierstrass +section tsums + +/-- If the terms in the sum `∑' (i : ι), F i` are uniformly bounded on `U` by a +summable function, and each term in the sum is differentiable on `U`, then so is the sum. -/ +lemma differentiable_on_tsum_of_summable_norm {u : ι → ℝ} + (hu : summable u) (hf : ∀ (i : ι), differentiable_on ℂ (F i) U) (hU : is_open U) + (hF_le : ∀ (i : ι) (w : ℂ), w ∈ U → ‖F i w‖ ≤ u i) : + differentiable_on ℂ (λ w : ℂ, ∑' (i : ι), F i w) U := +begin + classical, + have hc := (tendsto_uniformly_on_tsum hu hF_le).tendsto_locally_uniformly_on, + refine hc.differentiable_on (eventually_of_forall $ λ s, _) hU, + exact differentiable_on.sum (λ i hi, hf i), +end + +/-- If the terms in the sum `∑' (i : ι), F i` are uniformly bounded on `U` by a +summable function, then the sum of `deriv F i` at a point in `U` is the derivative of the +sum. -/ +lemma has_sum_deriv_of_summable_norm {u : ι → ℝ} + (hu : summable u) (hf : ∀ (i : ι), differentiable_on ℂ (F i) U) (hU : is_open U) + (hF_le : ∀ (i : ι) (w : ℂ), w ∈ U → ‖F i w‖ ≤ u i) (hz : z ∈ U) : + has_sum (λ (i : ι), deriv (F i) z) (deriv (λ w : ℂ, ∑' (i : ι), F i w) z) := +begin + rw has_sum, + have hc := (tendsto_uniformly_on_tsum hu hF_le).tendsto_locally_uniformly_on, + convert (hc.deriv (eventually_of_forall $ λ s, differentiable_on.sum + (λ i hi, hf i)) hU).tendsto_at hz using 1, + ext1 s, + exact (deriv_sum (λ i hi, (hf i).differentiable_at (hU.mem_nhds hz))).symm, +end + +end tsums + end complex diff --git a/src/analysis/complex/upper_half_plane/topology.lean b/src/analysis/complex/upper_half_plane/topology.lean index 418466831de71..9bbe36ca3bba2 100644 --- a/src/analysis/complex/upper_half_plane/topology.lean +++ b/src/analysis/complex/upper_half_plane/topology.lean @@ -9,6 +9,8 @@ import analysis.convex.normed import analysis.convex.complex import analysis.complex.re_im_topology import topology.homotopy.contractible +import geometry.manifold.mfderiv +import geometry.manifold.cont_mdiff_mfderiv /-! # Topology on the upper half plane @@ -19,7 +21,7 @@ various instances. noncomputable theory open set filter function topological_space complex -open_locale filter topology upper_half_plane +open_locale filter topology upper_half_plane manifold namespace upper_half_plane @@ -57,4 +59,18 @@ end instance : locally_compact_space ℍ := open_embedding_coe.locally_compact_space +instance upper_half_plane.charted_space : charted_space ℂ ℍ := +upper_half_plane.open_embedding_coe.singleton_charted_space + +instance upper_half_plane.smooth_manifold_with_corners : smooth_manifold_with_corners 𝓘(ℂ) ℍ := +upper_half_plane.open_embedding_coe.singleton_smooth_manifold_with_corners 𝓘(ℂ) + +/-- The inclusion map `ℍ → ℂ` is a smooth map of manifolds. -/ +lemma smooth_coe : smooth 𝓘(ℂ) 𝓘(ℂ) (coe : ℍ → ℂ) := +λ x, cont_mdiff_at_ext_chart_at + +/-- The inclusion map `ℍ → ℂ` is a differentiable map of manifolds. -/ +lemma mdifferentiable_coe : mdifferentiable 𝓘(ℂ) 𝓘(ℂ) (coe : ℍ → ℂ) := +smooth_coe.mdifferentiable + end upper_half_plane diff --git a/src/number_theory/modular_forms/basic.lean b/src/number_theory/modular_forms/basic.lean index ba025ac6d38c5..c83217272d5e5 100644 --- a/src/number_theory/modular_forms/basic.lean +++ b/src/number_theory/modular_forms/basic.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Birkbeck -/ -import geometry.manifold.mfderiv import analysis.complex.upper_half_plane.functions_bounded_at_infty import analysis.complex.upper_half_plane.topology import number_theory.modular_forms.slash_invariant_forms @@ -24,12 +23,6 @@ open_locale topology manifold upper_half_plane noncomputable theory -instance upper_half_plane.charted_space : charted_space ℂ ℍ := -upper_half_plane.open_embedding_coe.singleton_charted_space - -instance upper_half_plane.smooth_manifold_with_corners : smooth_manifold_with_corners 𝓘(ℂ) ℍ := -upper_half_plane.open_embedding_coe.singleton_smooth_manifold_with_corners 𝓘(ℂ) - local prefix `↑ₘ`:1024 := @coe _ (matrix (fin 2) (fin 2) _) _ local notation `GL(` n `, ` R `)`⁺ := matrix.GL_pos (fin n) R diff --git a/src/number_theory/modular_forms/jacobi_theta.lean b/src/number_theory/modular_forms/jacobi_theta.lean index 99e025059d26f..3711225ad9523 100644 --- a/src/number_theory/modular_forms/jacobi_theta.lean +++ b/src/number_theory/modular_forms/jacobi_theta.lean @@ -3,8 +3,10 @@ Copyright (c) 2023 David Loeffler. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: David Loeffler -/ -import analysis.complex.upper_half_plane.basic +import number_theory.modular_forms.basic import analysis.special_functions.gaussian +import analysis.calculus.series +import analysis.complex.locally_uniform_limit /-! # Jacobi's theta function @@ -13,43 +15,52 @@ This file defines the Jacobi theta function $$\theta(\tau) = \sum_{n \in \mathbb{Z}} \exp (i \pi n ^ 2 \tau),$$ and proves the modular transformation properties `θ (τ + 2) = θ τ` and -`θ (-1 / τ) = (-I * τ) ^ (1 / 2) * θ τ`, using Poisson's summation formula for the latter. +`θ (-1 / τ) = (-I * τ) ^ (1 / 2) * θ τ`, using Poisson's summation formula for the latter. We also +show that `θ` is differentiable on `ℍ`, and `θ(τ) - 1` has exponential decay as `im τ → ∞`. -/ -open complex real +open complex real asymptotics -open_locale real big_operators upper_half_plane +open_locale real big_operators upper_half_plane manifold /-- Jacobi's theta function `∑' (n : ℤ), exp (π * I * n ^ 2 * τ)`. -/ noncomputable def jacobi_theta (τ : ℍ) : ℂ := ∑' (n : ℤ), cexp (π * I * n ^ 2 * τ) -lemma jacobi_theta_unif_summable {R : ℝ} (hR : 0 < R) : +lemma norm_exp_mul_sq_le {z : ℂ} (hz : 0 < z.im) (n : ℤ) : + ‖cexp (π * I * n ^ 2 * z)‖ ≤ exp (-π * z.im) ^ n.nat_abs := +begin + let y := rexp (-π * z.im), + have h : y < 1, from exp_lt_one_iff.mpr (mul_neg_of_neg_of_pos (neg_lt_zero.mpr pi_pos) hz), + refine (le_of_eq _).trans (_ : y ^ (n ^ 2) ≤ _), + { rw [complex.norm_eq_abs, complex.abs_exp], + have : (↑π * I * n ^ 2 * z).re = (-π * z.im) * n ^ 2, + { rw [(by { push_cast, ring } : ↑π * I * n ^ 2 * z = ↑(π * n ^ 2) * (z * I)), + of_real_mul_re, mul_I_re], + ring }, + obtain ⟨m, hm⟩ := int.eq_coe_of_zero_le (sq_nonneg n), + rw [this, exp_mul, ←int.cast_pow, rpow_int_cast, hm, zpow_coe_nat] }, + { have : n ^ 2 = ↑(n.nat_abs ^ 2), by rw [nat.cast_pow, int.nat_abs_sq], + rw [this, zpow_coe_nat], + exact pow_le_pow_of_le_one (exp_pos _).le h.le ((sq n.nat_abs).symm ▸ n.nat_abs.le_mul_self) }, +end + +lemma exists_summable_bound_exp_mul_sq {R : ℝ} (hR : 0 < R) : ∃ (bd : ℤ → ℝ), (summable bd) ∧ - (∀ {τ : ℍ} (hτ : R ≤ τ.im) (n : ℤ), ‖cexp (π * I * n ^ 2 * τ)‖ ≤ bd n) := + (∀ {τ : ℂ} (hτ : R ≤ τ.im) (n : ℤ), ‖cexp (π * I * n ^ 2 * τ)‖ ≤ bd n) := begin let y := rexp (-π * R), have h : y < 1, from exp_lt_one_iff.mpr (mul_neg_of_neg_of_pos (neg_lt_zero.mpr pi_pos) hR), - refine ⟨λ n, y ^ |n|, summable_int_of_summable_nat _ _, λ τ hτ n, _⟩, swap 3, - { refine le_trans _ (_ : y ^ (n ^ 2) ≤ y ^ |n|), - { rw [complex.norm_eq_abs, complex.abs_exp], - have : (↑π * I * n ^ 2 * τ).re = (-π * (τ : ℂ).im) * n ^ 2, - { rw [(by { push_cast, ring } : ↑π * I * n ^ 2 * τ = ↑(π * n ^ 2) * (τ * I)), - of_real_mul_re, mul_I_re], - ring }, - obtain ⟨m, hm⟩ := int.eq_coe_of_zero_le (sq_nonneg n), - rw [this, exp_mul, ←int.cast_pow, rpow_int_cast, hm, zpow_coe_nat, zpow_coe_nat], - refine pow_le_pow_of_le_left (exp_pos _).le (real.exp_le_exp.mpr _) _, - rwa [mul_le_mul_left_of_neg (neg_lt_zero.mpr pi_pos), upper_half_plane.coe_im] }, - { rw [←inv_inv y, inv_zpow' _ (|n|), inv_zpow' _ (n ^ 2)], - refine zpow_le_of_le (one_le_inv (exp_pos _) h.le) (neg_le_neg _), - rw [int.abs_eq_nat_abs, ←int.nat_abs_sq, ←nat.cast_pow, nat.cast_le, sq], - exact n.nat_abs.le_mul_self } }, - all_goals { simp only [abs_neg, int.abs_coe_nat, zpow_coe_nat], - exact summable_geometric_of_lt_1 (real.exp_pos _).le h }, + refine ⟨λ n, y ^ n.nat_abs, summable_int_of_summable_nat _ _, λ τ hτ n, _⟩, swap 3, + { refine (norm_exp_mul_sq_le (hR.trans_le hτ) n).trans _, + refine pow_le_pow_of_le_left (exp_pos _).le (real.exp_le_exp.mpr _) _, + rwa [mul_le_mul_left_of_neg (neg_lt_zero.mpr pi_pos)] }, + all_goals { simpa only [int.nat_abs_neg, int.nat_abs_of_nat] + using summable_geometric_of_lt_1 (real.exp_pos _).le h }, end -lemma jacobi_theta_summable (τ : ℍ) : summable (λ n : ℤ, cexp (π * I * n ^ 2 * τ)) := -let ⟨bd, h, h'⟩ := jacobi_theta_unif_summable τ.im_pos in +lemma summable_exp_mul_sq {z : ℂ} (hz : 0 < z.im) : + summable (λ n : ℤ, cexp (π * I * n ^ 2 * z)) := +let ⟨bd, h, h'⟩ := exists_summable_bound_exp_mul_sq hz in summable_norm_iff.mp (summable_of_nonneg_of_le (λ n, norm_nonneg _) (h' $ le_refl _) h) lemma jacobi_theta_two_vadd (τ : ℍ) : jacobi_theta ((2 : ℝ) +ᵥ τ) = jacobi_theta τ := @@ -92,3 +103,80 @@ begin congr' 1, ring_nf } end + +lemma has_sum_nat_jacobi_theta (τ : ℍ) : + has_sum (λ (n : ℕ), cexp (π * I * (n + 1) ^ 2 * τ)) ((jacobi_theta τ - 1) / 2) := +begin + have := (summable_exp_mul_sq τ.im_pos).has_sum.sum_nat_of_sum_int, + rw ←@has_sum_nat_add_iff' ℂ _ _ _ _ 1 at this, + simp_rw [finset.sum_range_one, int.cast_neg, int.cast_coe_nat, nat.cast_zero, neg_zero, + int.cast_zero, sq (0:ℂ), mul_zero, zero_mul, neg_sq, ←mul_two, complex.exp_zero, + add_sub_assoc, (by norm_num : (1 : ℂ) - 1 * 2 = -1), ←sub_eq_add_neg, + nat.cast_add, nat.cast_one] at this, + convert this.div_const 2, + simp_rw mul_div_cancel _ two_ne_zero, +end + +lemma jacobi_theta_eq_tsum_nat (τ : ℍ) : + jacobi_theta τ = 1 + 2 * ∑' (n : ℕ), cexp (π * I * (n + 1) ^ 2 * τ) := +by rw [(has_sum_nat_jacobi_theta τ).tsum_eq, mul_div_cancel' _ (two_ne_zero' ℂ), ←add_sub_assoc, + add_sub_cancel'] + +/-- An explicit upper bound for `‖jacobi_theta τ - 1‖`. -/ +lemma norm_jacobi_theta_sub_one_le (τ : ℍ) : + ‖jacobi_theta τ - 1‖ ≤ 2 / (1 - exp (-π * τ.im)) * exp (-π * τ.im) := +begin + suffices : ‖∑' (n : ℕ), cexp (π * I * (n + 1) ^ 2 * τ)‖ ≤ exp (-π * τ.im) / (1 - exp (-π * τ.im)), + { calc ‖jacobi_theta τ - 1‖ = 2 * ‖∑' (n : ℕ), cexp (π * I * (n + 1) ^ 2 * τ)‖ : + by rw [sub_eq_iff_eq_add'.mpr (jacobi_theta_eq_tsum_nat τ), norm_mul, complex.norm_eq_abs, + complex.abs_two] + ... ≤ 2 * (rexp (-π * τ.im) / (1 - rexp (-π * τ.im))) : + by rwa [mul_le_mul_left (zero_lt_two' ℝ)] + ... = 2 / (1 - rexp (-π * τ.im)) * rexp (-π * τ.im) : by rw [div_mul_comm, mul_comm] }, + have : ∀ (n : ℕ), ‖cexp (π * I * (n + 1) ^ 2 * τ)‖ ≤ exp (-π * τ.im) ^ (n + 1), + { intro n, + simpa only [int.cast_add, int.cast_one] using norm_exp_mul_sq_le τ.im_pos (n + 1) }, + have s : has_sum (λ n : ℕ, rexp (-π * τ.im) ^ (n + 1)) (exp (-π * τ.im) / (1 - exp (-π * τ.im))), + { simp_rw [pow_succ, div_eq_mul_inv, has_sum_mul_left_iff (real.exp_ne_zero _)], + exact has_sum_geometric_of_lt_1 (exp_pos (-π * τ.im)).le + (exp_lt_one_iff.mpr $ (mul_neg_of_neg_of_pos (neg_lt_zero.mpr pi_pos) τ.im_pos)) }, + have aux : summable (λ (n : ℕ), ‖cexp (↑π * I * (↑n + 1) ^ 2 * ↑τ)‖), + from summable_of_nonneg_of_le (λ n, norm_nonneg _) this s.summable, + exact (norm_tsum_le_tsum_norm aux).trans + ((tsum_mono aux s.summable this).trans (le_of_eq s.tsum_eq)), +end + +/-- The norm of `jacobi_theta τ - 1` decays exponentially as `im τ → ∞`. -/ +lemma is_O_at_im_infty_jacobi_theta_sub_one : + is_O upper_half_plane.at_im_infty (λ τ, jacobi_theta τ - 1) (λ τ, rexp (-π * τ.im)) := +begin + simp_rw [is_O, is_O_with, filter.eventually, upper_half_plane.at_im_infty_mem], + refine ⟨2 / (1 - rexp (-π)), 1, (λ τ hτ, (norm_jacobi_theta_sub_one_le τ).trans _)⟩, + rw [real.norm_eq_abs, real.abs_exp], + refine mul_le_mul_of_nonneg_right _ (exp_pos _).le, + rw [div_le_div_left (zero_lt_two' ℝ), sub_le_sub_iff_left, exp_le_exp, neg_mul, neg_le_neg_iff], + { exact le_mul_of_one_le_right pi_pos.le hτ }, + { rw [sub_pos, exp_lt_one_iff, neg_mul, neg_lt_zero], exact mul_pos pi_pos τ.im_pos }, + { rw [sub_pos, exp_lt_one_iff, neg_lt_zero], exact pi_pos } +end + +lemma differentiable_at_tsum_exp_mul_sq (τ : ℍ) : + differentiable_at ℂ (λ z, ∑' (n : ℤ), cexp (π * I * n ^ 2 * z)) ↑τ := +begin + suffices : ∀ (y : ℝ) (hy : 0 < y), + differentiable_on ℂ (λ z, ∑' (n : ℤ), cexp (π * I * n ^ 2 * z)) {w : ℂ | y < im w}, + from let ⟨y, hy, hy'⟩ := exists_between τ.im_pos in (this y hy).differentiable_at + ((complex.continuous_im.is_open_preimage _ is_open_Ioi).mem_nhds (τ.coe_im ▸ hy')), + intros y hy, + have h1 : ∀ (n : ℤ) (w : ℂ) (hw : y < im w), differentiable_within_at ℂ + (λ (v : ℂ), cexp (↑π * I * ↑n ^ 2 * v)) {z : ℂ | y < im z} w, + from λ n w hw, (differentiable_at_id.const_mul _).cexp.differentiable_within_at, + have h2 : is_open {w : ℂ | y < im w}, from continuous_im.is_open_preimage _ is_open_Ioi, + obtain ⟨bd, bd_s, le_bd⟩ := exists_summable_bound_exp_mul_sq hy, + exact differentiable_on_tsum_of_summable_norm bd_s h1 h2 (λ i w hw, le_bd (le_of_lt hw) i), +end + +lemma mdifferentiable_jacobi_theta : mdifferentiable 𝓘(ℂ) 𝓘(ℂ) jacobi_theta := +λ τ, (differentiable_at_tsum_exp_mul_sq τ).mdifferentiable_at.comp τ τ.mdifferentiable_coe + +lemma continuous_jacobi_theta : continuous jacobi_theta := mdifferentiable_jacobi_theta.continuous From c941bb9426d62e266612b6d99e6c9fc93e7a1d07 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 30 Mar 2023 09:07:17 +0000 Subject: [PATCH 0734/1291] =?UTF-8?q?feat(data/set/finite):=20When=20`s=20?= =?UTF-8?q?=C3=97=CB=A2=20t`=20is=20finite=20(#18674)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The one non-trivial result is `infinite_image2`, because it requires only injectivity of the `f a` and `λ a, f a b` rather than of the uncurrying of `f`. --- archive/imo/imo2008_q2.lean | 2 +- src/data/set/finite.lean | 82 ++++++++++++++++++++++++++---- src/data/set/pointwise/finite.lean | 31 ++++++++++- 3 files changed, 103 insertions(+), 12 deletions(-) diff --git a/archive/imo/imo2008_q2.lean b/archive/imo/imo2008_q2.lean index 570dca94bad69..a20f897827eef 100644 --- a/archive/imo/imo2008_q2.lean +++ b/archive/imo/imo2008_q2.lean @@ -128,7 +128,7 @@ begin have hK_inf : set.infinite K, { intro h, apply hK_not_bdd, exact set.finite.bdd_above h }, - exact set.infinite_of_infinite_image g hK_inf }, + exact hK_inf.of_image g }, exact hW_inf.mono hW_sub_S, end diff --git a/src/data/set/finite.lean b/src/data/set/finite.lean index fe1ca45972561..0c891a6ca3202 100644 --- a/src/data/set/finite.lean +++ b/src/data/set/finite.lean @@ -117,8 +117,13 @@ protected def infinite (s : set α) : Prop := ¬ s.finite @[simp] lemma not_infinite {s : set α} : ¬ s.infinite ↔ s.finite := not_not +alias not_infinite ↔ _ finite.not_infinite + +attribute [simp] finite.not_infinite + /-- See also `finite_or_infinite`, `fintype_or_infinite`. -/ protected lemma finite_or_infinite (s : set α) : s.finite ∨ s.infinite := em _ +protected lemma infinite_or_finite (s : set α) : s.infinite ∨ s.finite := em' _ /-! ### Basic properties of `set.finite.to_finset` -/ @@ -595,6 +600,9 @@ h.bUnion hf @[simp] theorem finite_empty : (∅ : set α).finite := to_finite _ +protected lemma infinite.nonempty {s : set α} (h : s.infinite) : s.nonempty := +nonempty_iff_ne_empty.2 $ by { rintro rfl, exact h finite_empty } + @[simp] theorem finite_singleton (a : α) : ({a} : set α).finite := to_finite _ theorem finite_pure (a : α) : (pure a : set α).finite := to_finite _ @@ -639,17 +647,49 @@ lemma finite_lt_nat (n : ℕ) : set.finite {i | i < n} := to_finite _ lemma finite_le_nat (n : ℕ) : set.finite {i | i ≤ n} := to_finite _ -lemma finite.prod {s : set α} {t : set β} (hs : s.finite) (ht : t.finite) : - (s ×ˢ t : set (α × β)).finite := +section prod +variables {s : set α} {t : set β} + +protected lemma finite.prod (hs : s.finite) (ht : t.finite) : (s ×ˢ t : set (α × β)).finite := by { casesI hs, casesI ht, apply to_finite } -lemma finite.off_diag {s : set α} (hs : s.finite) : s.off_diag.finite := +lemma finite.of_prod_left (h : (s ×ˢ t : set (α × β)).finite) : t.nonempty → s.finite := +λ ⟨b, hb⟩, (h.image prod.fst).subset $ λ a ha, ⟨(a, b), ⟨ha, hb⟩, rfl⟩ + +lemma finite.of_prod_right (h : (s ×ˢ t : set (α × β)).finite) : s.nonempty → t.finite := +λ ⟨a, ha⟩, (h.image prod.snd).subset $ λ b hb, ⟨(a, b), ⟨ha, hb⟩, rfl⟩ + +protected lemma infinite.prod_left (hs : s.infinite) (ht : t.nonempty) : (s ×ˢ t).infinite := +λ h, hs $ h.of_prod_left ht + +protected lemma infinite.prod_right (ht : t.infinite) (hs : s.nonempty) : (s ×ˢ t).infinite := +λ h, ht $ h.of_prod_right hs + +protected lemma infinite_prod : + (s ×ˢ t).infinite ↔ s.infinite ∧ t.nonempty ∨ t.infinite ∧ s.nonempty := +begin + refine ⟨λ h, _, _⟩, + { simp_rw [set.infinite, and_comm (¬ _), ←not_imp], + by_contra', + exact h ((this.1 h.nonempty.snd).prod $ this.2 h.nonempty.fst) }, + { rintro (h | h), + { exact h.1.prod_left h.2 }, + { exact h.1.prod_right h.2 } } +end + +lemma finite_prod : (s ×ˢ t).finite ↔ (s.finite ∨ t = ∅) ∧ (t.finite ∨ s = ∅) := +by simp only [←not_infinite, set.infinite_prod, not_or_distrib, not_and_distrib, + not_nonempty_iff_eq_empty] + +protected lemma finite.off_diag (hs : s.finite) : s.off_diag.finite := by { classical, casesI hs, apply set.to_finite } -lemma finite.image2 (f : α → β → γ) {s : set α} {t : set β} (hs : s.finite) (ht : t.finite) : +protected lemma finite.image2 (f : α → β → γ) (hs : s.finite) (ht : t.finite) : (image2 f s t).finite := by { casesI hs, casesI ht, apply to_finite } +end prod + theorem finite.seq {f : set (α → β)} {s : set α} (hf : f.finite) (hs : s.finite) : (f.seq s).finite := by { classical, casesI hf, casesI hs, apply to_finite } @@ -944,9 +984,6 @@ lemma infinite.exists_subset_card_eq {s : set α} (hs : s.infinite) (n : ℕ) : ∃ t : finset α, ↑t ⊆ s ∧ t.card = n := ⟨((finset.range n).map (hs.nat_embedding _)).map (embedding.subtype _), by simp⟩ -lemma infinite.nonempty {s : set α} (h : s.infinite) : s.nonempty := -let a := infinite.nat_embedding s h 37 in ⟨a.1, a.2⟩ - lemma infinite_of_finite_compl [infinite α] {s : set α} (hs : sᶜ.finite) : s.infinite := λ h, set.infinite_univ (by simpa using hs.union h) @@ -962,14 +999,41 @@ lemma infinite.diff {s t : set α} (hs : s.infinite) (ht : t.finite) : (s \ t).i @[simp] lemma infinite_union {s t : set α} : (s ∪ t).infinite ↔ s.infinite ∨ t.infinite := by simp only [set.infinite, finite_union, not_and_distrib] -theorem infinite_of_infinite_image (f : α → β) {s : set α} (hs : (f '' s).infinite) : - s.infinite := +theorem infinite.of_image (f : α → β) {s : set α} (hs : (f '' s).infinite) : s.infinite := mt (finite.image f) hs theorem infinite_image_iff {s : set α} {f : α → β} (hi : inj_on f s) : (f '' s).infinite ↔ s.infinite := not_congr $ finite_image_iff hi +alias infinite_image_iff ↔ _ infinite.image + +attribute [protected] infinite.image + +section image2 +variables {f : α → β → γ} {s : set α} {t : set β} {a : α} {b : β} + +protected lemma infinite.image2_left (hs : s.infinite) (hb : b ∈ t) (hf : inj_on (λ a, f a b) s) : + (image2 f s t).infinite := +(hs.image hf).mono $ image_subset_image2_left hb + +protected lemma infinite.image2_right (ht : t.infinite) (ha : a ∈ s) (hf : inj_on (f a) t) : + (image2 f s t).infinite := +(ht.image hf).mono $ image_subset_image2_right ha + +theorem infinite_image2 (hfs : ∀ b ∈ t, inj_on (λ a, f a b) s) (hft : ∀ a ∈ s, inj_on (f a) t) : + (image2 f s t).infinite ↔ s.infinite ∧ t.nonempty ∨ t.infinite ∧ s.nonempty := +begin + refine ⟨λ h, set.infinite_prod.1 _, _⟩, + { rw ←image_uncurry_prod at h, + exact h.of_image _ }, + { rintro (⟨hs, b, hb⟩ | ⟨ht, a, ha⟩), + { exact hs.image2_left hb (hfs _ hb) }, + { exact ht.image2_right ha (hft _ ha) } } +end + +end image2 + theorem infinite_of_inj_on_maps_to {s : set α} {t : set β} {f : α → β} (hi : inj_on f s) (hm : maps_to f s t) (hs : s.infinite) : t.infinite := ((infinite_image_iff hi).2 hs).mono (maps_to'.mp hm) diff --git a/src/data/set/pointwise/finite.lean b/src/data/set/pointwise/finite.lean index 741a5f60cf39e..ba1fc40875cfd 100644 --- a/src/data/set/pointwise/finite.lean +++ b/src/data/set/pointwise/finite.lean @@ -6,10 +6,12 @@ Authors: Johan Commelin, Floris van Doorn import data.set.finite import data.set.pointwise.smul -/-! # Finiteness lemmas for pointwise operations on sets +/-! +# Finiteness lemmas for pointwise operations on sets > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. -> Any changes to this file require a corresponding PR to mathlib4.-/ +> Any changes to this file require a corresponding PR to mathlib4. +-/ open_locale pointwise @@ -74,6 +76,7 @@ section has_smul_set variables [has_smul α β] {s : set β} {a : α} @[to_additive] lemma finite.smul_set : s.finite → (a • s).finite := finite.image _ +@[to_additive] lemma infinite.of_smul_set : (a • s).infinite → s.infinite := infinite.of_image _ end has_smul_set @@ -85,6 +88,30 @@ lemma finite.vsub (hs : s.finite) (ht : t.finite) : set.finite (s -ᵥ t) := hs. end vsub +section cancel +variables [has_mul α] [is_left_cancel_mul α] [is_right_cancel_mul α] {s t : set α} + +@[to_additive] lemma infinite_mul : + (s * t).infinite ↔ s.infinite ∧ t.nonempty ∨ t.infinite ∧ s.nonempty := +infinite_image2 (λ _ _, (mul_left_injective _).inj_on _) (λ _ _, (mul_right_injective _).inj_on _) + +end cancel + +section group +variables [group α] [mul_action α β] {a : α} {s : set β} + +@[simp, to_additive] lemma finite_smul_set : (a • s).finite ↔ s.finite := +finite_image_iff $ (mul_action.injective _).inj_on _ + +@[simp, to_additive] lemma infinite_smul_set : (a • s).infinite ↔ s.infinite := +infinite_image_iff $ (mul_action.injective _).inj_on _ + +alias finite_smul_set ↔ finite.of_smul_set _ +alias infinite_smul_set ↔ _ infinite.smul_set + +attribute [to_additive] finite.of_smul_set infinite.smul_set + +end group end set open set From d97a0c9f7a7efe6d76d652c5a6b7c9c634b70e0a Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 30 Mar 2023 12:15:49 +0000 Subject: [PATCH 0735/1291] feat(linear_algebra/tensor_power): the tensor powers form a graded algebra (#10255) --- src/data/fin/tuple/basic.lean | 28 +++ .../tensor_algebra/to_tensor_power.lean | 163 +++++++++++++ src/linear_algebra/tensor_power.lean | 229 +++++++++++++++++- 3 files changed, 408 insertions(+), 12 deletions(-) create mode 100644 src/linear_algebra/tensor_algebra/to_tensor_power.lean diff --git a/src/data/fin/tuple/basic.lean b/src/data/fin/tuple/basic.lean index d2b4cb189fac6..eddc29b8946b3 100644 --- a/src/data/fin/tuple/basic.lean +++ b/src/data/fin/tuple/basic.lean @@ -310,6 +310,20 @@ begin simp [←nat_add_nat_add] }, end +/-- Appending a one-tuple to the left is the same as `fin.cons`. -/ +lemma append_left_eq_cons {α : Type*} {n : ℕ} (x₀ : fin 1 → α) (x : fin n → α): + fin.append x₀ x = fin.cons (x₀ 0) x ∘ fin.cast (add_comm _ _) := +begin + ext i, + refine fin.add_cases _ _ i; clear i, + { intro i, + rw [subsingleton.elim i 0, fin.append_left, function.comp_apply, eq_comm], + exact fin.cons_zero _ _, }, + { intro i, + rw [fin.append_right, function.comp_apply, fin.cast_nat_add, eq_comm, fin.add_nat_one], + exact fin.cons_succ _ _ _ }, +end + end append section repeat @@ -534,6 +548,20 @@ begin simp } end +/-- Appending a one-tuple to the right is the same as `fin.snoc`. -/ +lemma append_right_eq_snoc {α : Type*} {n : ℕ} (x : fin n → α) (x₀ : fin 1 → α) : + fin.append x x₀ = fin.snoc x (x₀ 0) := +begin + ext i, + refine fin.add_cases _ _ i; clear i, + { intro i, + rw [fin.append_left], + exact (@snoc_cast_succ _ (λ _, α) _ _ i).symm, }, + { intro i, + rw [subsingleton.elim i 0, fin.append_right], + exact (@snoc_last _ (λ _, α) _ _).symm, }, +end + lemma comp_init {α : Type*} {β : Type*} (g : α → β) (q : fin n.succ → α) : g ∘ (init q) = init (g ∘ q) := by { ext j, simp [init] } diff --git a/src/linear_algebra/tensor_algebra/to_tensor_power.lean b/src/linear_algebra/tensor_algebra/to_tensor_power.lean new file mode 100644 index 0000000000000..5d619d69bba0b --- /dev/null +++ b/src/linear_algebra/tensor_algebra/to_tensor_power.lean @@ -0,0 +1,163 @@ +/- +Copyright (c) 2021 Eric Wieser. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Eric Wieser +-/ +import linear_algebra.tensor_algebra.basic +import linear_algebra.tensor_power +/-! +# Tensor algebras as direct sums of tensor powers + +In this file we show that `tensor_algebra R M` is isomorphic to a direct sum of tensor powers, as +`tensor_algebra.equiv_direct_sum`. +-/ +open_locale direct_sum tensor_product + +variables {R M : Type*} [comm_semiring R] [add_comm_monoid M] [module R M] + +namespace tensor_power + +/-- The canonical embedding from a tensor power to the tensor algebra -/ +def to_tensor_algebra {n} : ⨂[R]^n M →ₗ[R] tensor_algebra R M := +pi_tensor_product.lift (tensor_algebra.tprod R M n) + +@[simp] +lemma to_tensor_algebra_tprod {n} (x : fin n → M) : + tensor_power.to_tensor_algebra (pi_tensor_product.tprod R x) = tensor_algebra.tprod R M n x := +pi_tensor_product.lift.tprod _ + +@[simp] +lemma to_tensor_algebra_ghas_one : + (@graded_monoid.ghas_one.one _ (λ n, ⨂[R]^n M) _ _).to_tensor_algebra = 1 := +tensor_power.to_tensor_algebra_tprod _ + +@[simp] +lemma to_tensor_algebra_ghas_mul {i j} (a : ⨂[R]^i M) (b : ⨂[R]^j M) : + (@graded_monoid.ghas_mul.mul _ (λ n, ⨂[R]^n M) _ _ _ _ a b).to_tensor_algebra + = a.to_tensor_algebra * b.to_tensor_algebra := +begin + -- change `a` and `b` to `tprod R a` and `tprod R b` + rw [tensor_power.ghas_mul_eq_coe_linear_map, ←linear_map.compr₂_apply, + ←@linear_map.mul_apply' R, ←linear_map.compl₂_apply, ←linear_map.comp_apply], + refine linear_map.congr_fun (linear_map.congr_fun _ a) b, + clear a b, + ext a b, + simp only [linear_map.compr₂_apply, linear_map.mul_apply', + linear_map.compl₂_apply, linear_map.comp_apply, linear_map.comp_multilinear_map_apply, + pi_tensor_product.lift.tprod, tensor_power.tprod_mul_tprod, + tensor_power.to_tensor_algebra_tprod, tensor_algebra.tprod_apply, ←ghas_mul_eq_coe_linear_map], + refine eq.trans _ list.prod_append, + congr', + rw [←list.map_of_fn _ (tensor_algebra.ι R), ←list.map_of_fn _ (tensor_algebra.ι R), + ←list.map_of_fn _ (tensor_algebra.ι R), ←list.map_append, list.of_fn_fin_append], +end + +@[simp] +lemma to_tensor_algebra_galgebra_to_fun (r : R) : + (@direct_sum.galgebra.to_fun _ R (λ n, ⨂[R]^n M) _ _ _ _ _ _ _ r).to_tensor_algebra + = algebra_map _ _ r := +by rw [tensor_power.galgebra_to_fun_def, tensor_power.algebra_map₀_eq_smul_one, linear_map.map_smul, + tensor_power.to_tensor_algebra_ghas_one, algebra.algebra_map_eq_smul_one] + +end tensor_power + +namespace tensor_algebra + +/-- The canonical map from a direct sum of tensor powers to the tensor algebra. -/ +def of_direct_sum : (⨁ n, ⨂[R]^n M) →ₐ[R] tensor_algebra R M := +direct_sum.to_algebra _ _ (λ n, tensor_power.to_tensor_algebra) + tensor_power.to_tensor_algebra_ghas_one + (λ i j, tensor_power.to_tensor_algebra_ghas_mul) + (tensor_power.to_tensor_algebra_galgebra_to_fun) + +@[simp] lemma of_direct_sum_of_tprod {n} (x : fin n → M) : + of_direct_sum (direct_sum.of _ n (pi_tensor_product.tprod R x)) = tprod R M n x := +(direct_sum.to_add_monoid_of _ _ _).trans (tensor_power.to_tensor_algebra_tprod _) + +/-- The canonical map from the tensor algebra to a direct sum of tensor powers. -/ +def to_direct_sum : tensor_algebra R M →ₐ[R] ⨁ n, ⨂[R]^n M := +tensor_algebra.lift R $ + direct_sum.lof R ℕ (λ n, ⨂[R]^n M) _ ∘ₗ + (linear_equiv.symm $ pi_tensor_product.subsingleton_equiv (0 : fin 1) : M ≃ₗ[R] _).to_linear_map + +@[simp] lemma to_direct_sum_ι (x : M) : + to_direct_sum (ι R x) = + direct_sum.of (λ n, ⨂[R]^n M) _ (pi_tensor_product.tprod R (λ _ : fin 1, x)) := +tensor_algebra.lift_ι_apply _ _ + +lemma of_direct_sum_comp_to_direct_sum : + of_direct_sum.comp to_direct_sum = alg_hom.id R (tensor_algebra R M) := +begin + ext, + simp [direct_sum.lof_eq_of, tprod_apply], +end + +@[simp] lemma of_direct_sum_to_direct_sum (x : tensor_algebra R M) : + of_direct_sum x.to_direct_sum = x := +alg_hom.congr_fun of_direct_sum_comp_to_direct_sum x + +@[simp] lemma mk_reindex_cast {n m : ℕ} (h : n = m) (x : ⨂[R]^n M) : + graded_monoid.mk m (pi_tensor_product.reindex R M (equiv.cast $ congr_arg fin h) x) = + graded_monoid.mk n x := +eq.symm (pi_tensor_product.graded_monoid_eq_of_reindex_cast h rfl) + +@[simp] lemma mk_reindex_fin_cast {n m : ℕ} (h : n = m) (x : ⨂[R]^n M) : + graded_monoid.mk m (pi_tensor_product.reindex R M (fin.cast h).to_equiv x) = + graded_monoid.mk n x := +by rw [fin.cast_to_equiv, mk_reindex_cast h] + +/-- The product of tensor products made of a single vector is the same as a single product of +all the vectors. -/ +lemma _root_.tensor_power.list_prod_graded_monoid_mk_single (n : ℕ) (x : fin n → M) : + ((list.fin_range n).map + (λ a, (graded_monoid.mk _ (pi_tensor_product.tprod R (λ i : fin 1, x a)) + : graded_monoid (λ n, ⨂[R]^n M)))).prod = + graded_monoid.mk n (pi_tensor_product.tprod R x) := +begin + refine fin.cons_induction _ _ x; clear x, + { rw [list.fin_range_zero, list.map_nil, list.prod_nil], + refl, }, + { intros n x₀ x ih, + rw [list.fin_range_succ_eq_map, list.map_cons, list.prod_cons, list.map_map, function.comp], + simp_rw [fin.cons_zero, fin.cons_succ], + rw [ih, graded_monoid.mk_mul_mk, tensor_power.tprod_mul_tprod], + refine tensor_power.graded_monoid_eq_of_cast (add_comm _ _) _, + dsimp only [graded_monoid.mk], + rw [tensor_power.cast_tprod], + simp_rw [fin.append_left_eq_cons, function.comp], + congr' 1 with i, + congr' 1, + rw [fin.cast_trans, fin.cast_refl, order_iso.refl_apply] }, +end + +lemma to_direct_sum_tensor_power_tprod {n} (x : fin n → M) : + to_direct_sum (tprod R M n x) = direct_sum.of _ n (pi_tensor_product.tprod R x) := +begin + rw [tprod_apply, alg_hom.map_list_prod, list.map_of_fn, function.comp], + simp_rw to_direct_sum_ι, + dsimp only, + rw direct_sum.list_prod_of_fn_of_eq_dprod, + apply direct_sum.of_eq_of_graded_monoid_eq, + rw graded_monoid.mk_list_dprod, + rw tensor_power.list_prod_graded_monoid_mk_single, +end + +lemma to_direct_sum_comp_of_direct_sum : + to_direct_sum.comp of_direct_sum = alg_hom.id R (⨁ n, ⨂[R]^n M) := +begin + ext, + simp [direct_sum.lof_eq_of, -tprod_apply, to_direct_sum_tensor_power_tprod], +end + +@[simp] lemma to_direct_sum_of_direct_sum (x : ⨁ n, ⨂[R]^n M) : + (of_direct_sum x).to_direct_sum = x := +alg_hom.congr_fun to_direct_sum_comp_of_direct_sum x + +/-- The tensor algebra is isomorphic to a direct sum of tensor powers. -/ +@[simps] +def equiv_direct_sum : tensor_algebra R M ≃ₐ[R] ⨁ n, ⨂[R]^n M := +alg_equiv.of_alg_hom to_direct_sum of_direct_sum + to_direct_sum_comp_of_direct_sum + of_direct_sum_comp_to_direct_sum + +end tensor_algebra diff --git a/src/linear_algebra/tensor_power.lean b/src/linear_algebra/tensor_power.lean index fe707662fff8c..828dda9d634fb 100644 --- a/src/linear_algebra/tensor_power.lean +++ b/src/linear_algebra/tensor_power.lean @@ -6,7 +6,7 @@ Authors: Eric Wieser import linear_algebra.pi_tensor_product import logic.equiv.fin -import algebra.graded_monoid +import algebra.direct_sum.algebra /-! # Tensor power of a semimodule over a commutative semirings @@ -19,13 +19,8 @@ abbreviation for `⨂[R] i : fin n, M`. ## Main definitions: -* `tensor_power.ghas_one` -* `tensor_power.ghas_mul` - -## TODO - -Show `direct_sum.galgebra R (λ i, ⨂[R]^i M)` and `algebra R (⨁ n : ℕ, ⨂[R]^n M)`. - +* `tensor_power.gsemiring`: the tensor powers form a graded semiring. +* `tensor_power.galgebra`: the tensor powers form a graded algebra. ## Implementation notes @@ -46,17 +41,34 @@ variables {R : Type*} {M : Type*} [comm_semiring R] [add_comm_monoid M] [module localized "notation (name := tensor_power) `⨂[`:100 R `]^`:80 n:max := tensor_power R n" in tensor_product +namespace pi_tensor_product + +/-- Two dependent pairs of tensor products are equal if their index is equal and the contents +are equal after a canonical reindexing. -/ +@[ext] +lemma graded_monoid_eq_of_reindex_cast {ιι : Type*} {ι : ιι → Type*} + [dι : Π ii, decidable_eq (ι ii)] : + ∀ {a b : graded_monoid (λ ii, ⨂[R] i : ι ii, M)} (h : a.fst = b.fst), + reindex R M (equiv.cast $ congr_arg ι h) a.snd = b.snd → a = b +| ⟨ai, a⟩ ⟨bi, b⟩ := λ (hi : ai = bi) (h : reindex R M _ a = b), +begin + subst hi, + simpa using h, +end + +end pi_tensor_product + namespace tensor_power -open_locale tensor_product +open_locale tensor_product direct_sum open pi_tensor_product /-- As a graded monoid, `⨂[R]^i M` has a `1 : ⨂[R]^0 M`. -/ instance ghas_one : graded_monoid.ghas_one (λ i, ⨂[R]^i M) := -{ one := tprod R fin.elim0 } +{ one := tprod R $ @fin.elim0' M } local notation `ₜ1` := @graded_monoid.ghas_one.one ℕ (λ i, ⨂[R]^i M) _ _ -lemma ghas_one_def : ₜ1 = tprod R fin.elim0 := rfl +lemma ghas_one_def : ₜ1 = tprod R (@fin.elim0' M) := rfl /-- A variant of `pi_tensor_prod.tmul_equiv` with the result indexed by `fin (n + m)`. -/ def mul_equiv {n m : ℕ} : (⨂[R]^n M) ⊗[R] (⨂[R]^m M) ≃ₗ[R] ⨂[R]^(n + m) M := @@ -64,10 +76,203 @@ def mul_equiv {n m : ℕ} : (⨂[R]^n M) ⊗[R] (⨂[R]^m M) ≃ₗ[R] ⨂[R]^(n /-- As a graded monoid, `⨂[R]^i M` has a `(*) : ⨂[R]^i M → ⨂[R]^j M → ⨂[R]^(i + j) M`. -/ instance ghas_mul : graded_monoid.ghas_mul (λ i, ⨂[R]^i M) := -{ mul := λ i j a b, mul_equiv (a ⊗ₜ b) } +{ mul := λ i j a b, (tensor_product.mk R _ _).compr₂ ↑(mul_equiv : _ ≃ₗ[R] ⨂[R]^(i + j) M) a b} local infix ` ₜ* `:70 := @graded_monoid.ghas_mul.mul ℕ (λ i, ⨂[R]^i M) _ _ _ _ lemma ghas_mul_def {i j} (a : ⨂[R]^i M) (b : ⨂[R]^j M) : a ₜ* b = mul_equiv (a ⊗ₜ b) := rfl +lemma ghas_mul_eq_coe_linear_map {i j} (a : ⨂[R]^i M) (b : ⨂[R]^j M) : + a ₜ* b = + ((tensor_product.mk R _ _).compr₂ ↑(mul_equiv : _ ≃ₗ[R] ⨂[R]^(i + j) M) + : ⨂[R]^i M →ₗ[R] ⨂[R]^j M →ₗ[R] ⨂[R]^(i + j) M) a b := rfl + +variables (R M) + +/-- Cast between "equal" tensor powers. -/ +def cast {i j} (h : i = j) : ⨂[R]^i M ≃ₗ[R] (⨂[R]^j M) := +reindex R M (fin.cast h).to_equiv + +lemma cast_tprod {i j} (h : i = j) (a : fin i → M) : + cast R M h (tprod R a) = tprod R (a ∘ fin.cast h.symm) := +reindex_tprod _ _ + +@[simp] lemma cast_refl {i} (h : i = i) : cast R M h = linear_equiv.refl _ _ := +(congr_arg (λ f, reindex R M (rel_iso.to_equiv f)) $ fin.cast_refl h).trans reindex_refl + +@[simp] lemma cast_symm {i j} (h : i = j) : (cast R M h).symm = cast R M h.symm := reindex_symm _ + +@[simp] lemma cast_trans {i j k} (h : i = j) (h' : j = k) : + (cast R M h).trans (cast R M h') = cast R M (h.trans h') := reindex_trans _ _ + +variables {R M} + +@[simp] lemma cast_cast {i j k} (h : i = j) (h' : j = k) (a : ⨂[R]^i M) : + cast R M h' (cast R M h a) = cast R M (h.trans h') a := reindex_reindex _ _ _ + +@[ext] +lemma graded_monoid_eq_of_cast {a b : graded_monoid (λ n, ⨂[R] i : fin n, M)} + (h : a.fst = b.fst) (h2 : cast R M h a.snd = b.snd) : a = b := +begin + refine graded_monoid_eq_of_reindex_cast h _, + rw cast at h2, + rw [←fin.cast_to_equiv, ← h2], +end + +-- named to match `fin.cast_eq_cast` +lemma cast_eq_cast {i j} (h : i = j) : ⇑(cast R M h) = _root_.cast (congr_arg _ h) := +begin + subst h, + rw [cast_refl], + refl, +end + +variables (R) +include R +lemma tprod_mul_tprod {na nb} (a : fin na → M) (b : fin nb → M) : + tprod R a ₜ* tprod R b = tprod R (fin.append a b) := +begin + dsimp [ghas_mul_def, mul_equiv], + rw [tmul_equiv_apply R M a b], + refine (reindex_tprod _ _).trans _, + congr' 1, + dsimp only [fin.append, fin_sum_fin_equiv, equiv.coe_fn_symm_mk], + apply funext, + apply fin.add_cases; simp, +end +omit R +variables {R} + +lemma one_mul {n} (a : ⨂[R]^n M) : + cast R M (zero_add n) (ₜ1 ₜ* a) = a := +begin + rw [ghas_mul_def, ghas_one_def], + induction a using pi_tensor_product.induction_on with r a x y hx hy, + { dsimp only at a, + rw [tensor_product.tmul_smul, linear_equiv.map_smul, linear_equiv.map_smul, ←ghas_mul_def, + tprod_mul_tprod, cast_tprod], + congr' 2 with i, + rw fin.elim0'_append, + refine congr_arg a (fin.ext _), + simp }, + { rw [tensor_product.tmul_add, map_add, map_add, hx, hy], }, +end + +lemma mul_one {n} (a : ⨂[R]^n M) : cast R M (add_zero _) (a ₜ* ₜ1) = a := +begin + rw [ghas_mul_def, ghas_one_def], + induction a using pi_tensor_product.induction_on with r a x y hx hy, + { dsimp only at a, + rw [←tensor_product.smul_tmul', linear_equiv.map_smul, linear_equiv.map_smul, ←ghas_mul_def, + tprod_mul_tprod R a _, cast_tprod], + congr' 2 with i, + rw fin.append_elim0', + refine congr_arg a (fin.ext _), + simp }, + { rw [tensor_product.add_tmul, map_add, map_add, hx, hy], }, +end + +lemma mul_assoc {na nb nc} (a : ⨂[R]^na M) (b : ⨂[R]^nb M) (c : ⨂[R]^nc M) : + cast R M (add_assoc _ _ _) ((a ₜ* b) ₜ* c) = a ₜ* (b ₜ* c) := +begin + let mul : Π (n m : ℕ), (⨂[R]^n M) →ₗ[R] (⨂[R]^m M) →ₗ[R] ⨂[R]^(n + m) M := + (λ n m, (tensor_product.mk R _ _).compr₂ ↑(mul_equiv : _ ≃ₗ[R] ⨂[R]^(n + m) M)), + -- replace `a`, `b`, `c` with `tprod R a`, `tprod R b`, `tprod R c` + let e : ⨂[R]^(na + nb + nc) M ≃ₗ[R] ⨂[R]^(na + (nb + nc)) M := cast R M (add_assoc _ _ _), + let lhs : (⨂[R]^na M) →ₗ[R] (⨂[R]^nb M) →ₗ[R] (⨂[R]^nc M) →ₗ[R] (⨂[R]^(na + (nb + nc)) M) := + (linear_map.llcomp R _ _ _ ((mul _ nc).compr₂ e.to_linear_map)).comp + (mul na nb), + have lhs_eq : ∀ a b c, lhs a b c = e ((a ₜ* b) ₜ* c) := λ _ _ _, rfl, + let rhs : (⨂[R]^na M) →ₗ[R] (⨂[R]^nb M) →ₗ[R] (⨂[R]^nc M) →ₗ[R] (⨂[R]^(na + (nb + nc)) M) := + (linear_map.llcomp R _ _ _ (linear_map.lflip R _ _ _) $ + (linear_map.llcomp R _ _ _ (mul na _).flip).comp (mul nb nc)).flip, + have rhs_eq : ∀ a b c, rhs a b c = (a ₜ* (b ₜ* c)) := λ _ _ _, rfl, + suffices : lhs = rhs, + from linear_map.congr_fun (linear_map.congr_fun (linear_map.congr_fun this a) b) c, + ext a b c, + -- clean up + simp only [linear_map.comp_multilinear_map_apply, lhs_eq, rhs_eq, tprod_mul_tprod, e, + cast_tprod], + congr' with j, + rw fin.append_assoc, + refine congr_arg (fin.append a (fin.append b c)) (fin.ext _), + rw [fin.coe_cast, fin.coe_cast], +end + +-- for now we just use the default for the `gnpow` field as it's easier. +instance gmonoid : graded_monoid.gmonoid (λ i, ⨂[R]^i M) := +{ one_mul := λ a, graded_monoid_eq_of_cast (zero_add _) (one_mul _), + mul_one := λ a, graded_monoid_eq_of_cast (add_zero _) (mul_one _), + mul_assoc := λ a b c, graded_monoid_eq_of_cast (add_assoc _ _ _) (mul_assoc _ _ _), + ..tensor_power.ghas_mul, + ..tensor_power.ghas_one, } + +/-- The canonical map from `R` to `⨂[R]^0 M` corresponding to the algebra_map of the tensor +algebra. -/ +def algebra_map₀ : R ≃ₗ[R] ⨂[R]^0 M := +linear_equiv.symm $ is_empty_equiv (fin 0) + +lemma algebra_map₀_eq_smul_one (r : R) : + (algebra_map₀ r : ⨂[R]^0 M) = r • ₜ1 := +by { simp [algebra_map₀], congr } + +lemma algebra_map₀_one : (algebra_map₀ 1 : ⨂[R]^0 M) = ₜ1 := +(algebra_map₀_eq_smul_one 1).trans (one_smul _ _) + +lemma algebra_map₀_mul {n} (r : R) (a : ⨂[R]^n M) : + cast R M (zero_add _) (algebra_map₀ r ₜ* a) = r • a := +by rw [ghas_mul_eq_coe_linear_map, algebra_map₀_eq_smul_one, linear_map.map_smul₂, + linear_equiv.map_smul, ←ghas_mul_eq_coe_linear_map, one_mul] + +lemma mul_algebra_map₀ {n} (r : R) (a : ⨂[R]^n M) : + cast R M (add_zero _) (a ₜ* algebra_map₀ r) = r • a := +by rw [ghas_mul_eq_coe_linear_map, algebra_map₀_eq_smul_one, linear_map.map_smul, + linear_equiv.map_smul, ←ghas_mul_eq_coe_linear_map, mul_one] + +lemma algebra_map₀_mul_algebra_map₀ (r s : R) : + cast R M (add_zero _) (algebra_map₀ r ₜ* algebra_map₀ s) = algebra_map₀ (r * s) := +begin + rw [←smul_eq_mul, linear_equiv.map_smul], + exact algebra_map₀_mul r (@algebra_map₀ R M _ _ _ s), +end + +instance gsemiring : direct_sum.gsemiring (λ i, ⨂[R]^i M) := +{ mul_zero := λ i j a, linear_map.map_zero _, + zero_mul := λ i j b, linear_map.map_zero₂ _ _, + mul_add := λ i j a b₁ b₂, linear_map.map_add _ _ _, + add_mul := λ i j a₁ a₂ b, linear_map.map_add₂ _ _ _ _, + nat_cast := λ n, algebra_map₀ (n : R), + nat_cast_zero := by rw [nat.cast_zero, map_zero], + nat_cast_succ := λ n, by rw [nat.cast_succ, map_add, algebra_map₀_one], + ..tensor_power.gmonoid } + +example : semiring (⨁ n : ℕ, ⨂[R]^n M) := by apply_instance + +/-- The tensor powers form a graded algebra. + +Note that this instance implies `algebra R (⨁ n : ℕ, ⨂[R]^n M)` via `direct_sum.algebra`. -/ +instance galgebra : direct_sum.galgebra R (λ i, ⨂[R]^i M) := +{ to_fun := (algebra_map₀ : R ≃ₗ[R] ⨂[R]^0 M).to_linear_map.to_add_monoid_hom, + map_one := algebra_map₀_one, + map_mul := λ r s, graded_monoid_eq_of_cast rfl begin + rw [←linear_equiv.eq_symm_apply], + have := algebra_map₀_mul_algebra_map₀ r s, + exact this.symm, + end, + commutes := λ r x, graded_monoid_eq_of_cast (add_comm _ _) begin + have := (algebra_map₀_mul r x.snd).trans (mul_algebra_map₀ r x.snd).symm, + rw [←linear_equiv.eq_symm_apply, cast_symm], + rw [←linear_equiv.eq_symm_apply, cast_symm, cast_cast] at this, + exact this, + end, + smul_def := λ r x, graded_monoid_eq_of_cast (zero_add x.fst).symm begin + rw [←linear_equiv.eq_symm_apply, cast_symm], + exact (algebra_map₀_mul r x.snd).symm, + end } + +lemma galgebra_to_fun_def (r : R) : + @direct_sum.galgebra.to_fun ℕ R (λ i, ⨂[R]^i M) _ _ _ _ _ _ _ r = algebra_map₀ r := rfl + +example : algebra R (⨁ n : ℕ, ⨂[R]^n M) := by apply_instance + end tensor_power From ec1c7d810034d4202b0dd239112d1792be9f6fdc Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Thu, 30 Mar 2023 17:06:28 +0000 Subject: [PATCH 0736/1291] chore(ring_theory/ideal/local_ring): move ring_equiv.local_ring into local_ring.lean (#18686) By having `ring_theory.ideal.local_ring` depend on `logic.equiv.transfer_instance`, rather than the other way around, we avoid importing all the dependencies for `local_ring` into files that just need to transfer "elementary" algebraic structures across equivalences. Co-authored-by: Scott Morrison --- src/algebra/module/injective.lean | 1 + src/algebraic_topology/dold_kan/degeneracies.lean | 1 + src/logic/equiv/transfer_instance.lean | 14 +------------- src/ring_theory/ideal/local_ring.lean | 12 ++++++++++++ 4 files changed, 15 insertions(+), 13 deletions(-) diff --git a/src/algebra/module/injective.lean b/src/algebra/module/injective.lean index e41135936b0af..b176046e376a0 100644 --- a/src/algebra/module/injective.lean +++ b/src/algebra/module/injective.lean @@ -6,6 +6,7 @@ Authors: Jujian Zhang import category_theory.preadditive.injective import ring_theory.ideal.basic +import linear_algebra.linear_pmap /-! # Injective modules diff --git a/src/algebraic_topology/dold_kan/degeneracies.lean b/src/algebraic_topology/dold_kan/degeneracies.lean index 3a831d4e85ee4..35aac95468ce1 100644 --- a/src/algebraic_topology/dold_kan/degeneracies.lean +++ b/src/algebraic_topology/dold_kan/degeneracies.lean @@ -5,6 +5,7 @@ Authors: Joël Riou -/ import algebraic_topology.dold_kan.decomposition +import tactic.fin_cases /-! diff --git a/src/logic/equiv/transfer_instance.lean b/src/logic/equiv/transfer_instance.lean index 16a3667c8ae7c..1db07027c0132 100644 --- a/src/logic/equiv/transfer_instance.lean +++ b/src/logic/equiv/transfer_instance.lean @@ -3,10 +3,9 @@ Copyright (c) 2018 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl -/ -import algebra.algebra.basic +import algebra.algebra.equiv import algebra.field.basic import logic.equiv.defs -import ring_theory.ideal.local_ring /-! # Transfer algebraic structures across `equiv`s @@ -409,14 +408,3 @@ end R end instances end equiv - -namespace ring_equiv - -@[reducible] protected lemma local_ring {A B : Type*} [comm_semiring A] [local_ring A] - [comm_semiring B] (e : A ≃+* B) : local_ring B := -begin - haveI := e.symm.to_equiv.nontrivial, - exact local_ring.of_surjective (e : A →+* B) e.surjective -end - -end ring_equiv diff --git a/src/ring_theory/ideal/local_ring.lean b/src/ring_theory/ideal/local_ring.lean index c3d77a19bdc4c..887c846030036 100644 --- a/src/ring_theory/ideal/local_ring.lean +++ b/src/ring_theory/ideal/local_ring.lean @@ -7,6 +7,7 @@ Authors: Kenny Lau, Chris Hughes, Mario Carneiro import algebra.algebra.basic import ring_theory.ideal.operations import ring_theory.jacobson_ideal +import logic.equiv.transfer_instance /-! @@ -460,3 +461,14 @@ end field lemma local_ring.maximal_ideal_eq_bot {R : Type*} [field R] : local_ring.maximal_ideal R = ⊥ := local_ring.is_field_iff_maximal_ideal_eq.mp (field.to_is_field R) + +namespace ring_equiv + +@[reducible] protected lemma local_ring {A B : Type*} [comm_semiring A] [local_ring A] + [comm_semiring B] (e : A ≃+* B) : local_ring B := +begin + haveI := e.symm.to_equiv.nontrivial, + exact local_ring.of_surjective (e : A →+* B) e.surjective +end + +end ring_equiv From f8d8465c3c392a93b9ed226956e26dee00975946 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Thu, 30 Mar 2023 17:06:29 +0000 Subject: [PATCH 0737/1291] chore(category_theory/preadditive/projective): split out two lemmas to reduce imports (#18688) Motivated by all the unneeded imports visible in https://tqft.net/mathlib4/2023-03-29/category_theory.monoidal.tor.pdf Co-authored-by: Scott Morrison --- src/algebra/module/injective.lean | 1 + src/category_theory/abelian/injective.lean | 1 + src/category_theory/abelian/projective.lean | 1 + .../preadditive/injective.lean | 31 ----------- src/category_theory/preadditive/opposite.lean | 1 - .../preadditive/projective.lean | 32 +---------- .../preadditive/yoneda/injective.lean | 53 +++++++++++++++++++ .../preadditive/yoneda/projective.lean | 53 +++++++++++++++++++ 8 files changed, 110 insertions(+), 63 deletions(-) create mode 100644 src/category_theory/preadditive/yoneda/injective.lean create mode 100644 src/category_theory/preadditive/yoneda/projective.lean diff --git a/src/algebra/module/injective.lean b/src/algebra/module/injective.lean index b176046e376a0..43b956e32afd4 100644 --- a/src/algebra/module/injective.lean +++ b/src/algebra/module/injective.lean @@ -5,6 +5,7 @@ Authors: Jujian Zhang -/ import category_theory.preadditive.injective +import algebra.category.Module.epi_mono import ring_theory.ideal.basic import linear_algebra.linear_pmap diff --git a/src/category_theory/abelian/injective.lean b/src/category_theory/abelian/injective.lean index 65665d046789e..bf89fc5c7b7ca 100644 --- a/src/category_theory/abelian/injective.lean +++ b/src/category_theory/abelian/injective.lean @@ -7,6 +7,7 @@ Authors: Jakob von Raumer import category_theory.abelian.exact import category_theory.preadditive.injective import category_theory.preadditive.yoneda.limits +import category_theory.preadditive.yoneda.injective /-! # Injective objects in abelian categories diff --git a/src/category_theory/abelian/projective.lean b/src/category_theory/abelian/projective.lean index 0ce42bf45f3bb..0868de8756aaa 100644 --- a/src/category_theory/abelian/projective.lean +++ b/src/category_theory/abelian/projective.lean @@ -7,6 +7,7 @@ import algebra.homology.quasi_iso import category_theory.abelian.homology import category_theory.preadditive.projective_resolution import category_theory.preadditive.yoneda.limits +import category_theory.preadditive.yoneda.projective /-! # Abelian categories with enough projectives have projective resolutions diff --git a/src/category_theory/preadditive/injective.lean b/src/category_theory/preadditive/injective.lean index 25382e26823db..b61b35c1b5b72 100644 --- a/src/category_theory/preadditive/injective.lean +++ b/src/category_theory/preadditive/injective.lean @@ -3,11 +3,7 @@ Copyright (c) 2022 Jujian Zhang. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jujian Zhang, Kevin Buzzard -/ - -import algebra.homology.exact -import category_theory.types import category_theory.preadditive.projective -import category_theory.limits.shapes.biproducts /-! # Injective objects and categories with enough injectives @@ -188,33 +184,6 @@ lemma injective_of_adjoint (adj : L ⊣ R) (J : D) [injective J] : injective $ R end adjunction -section preadditive -variables [preadditive C] - -lemma injective_iff_preserves_epimorphisms_preadditive_yoneda_obj (J : C) : - injective J ↔ (preadditive_yoneda.obj J).preserves_epimorphisms := -begin - rw injective_iff_preserves_epimorphisms_yoneda_obj, - refine ⟨λ (h : (preadditive_yoneda.obj J ⋙ (forget _)).preserves_epimorphisms), _, _⟩, - { exactI functor.preserves_epimorphisms_of_preserves_of_reflects (preadditive_yoneda.obj J) - (forget _) }, - { introI, - exact (infer_instance : (preadditive_yoneda.obj J ⋙ forget _).preserves_epimorphisms) } -end - -lemma injective_iff_preserves_epimorphisms_preadditive_yoneda_obj' (J : C) : - injective J ↔ (preadditive_yoneda_obj J).preserves_epimorphisms := -begin - rw injective_iff_preserves_epimorphisms_yoneda_obj, - refine ⟨λ (h : (preadditive_yoneda_obj J ⋙ (forget _)).preserves_epimorphisms), _, _⟩, - { exactI functor.preserves_epimorphisms_of_preserves_of_reflects (preadditive_yoneda_obj J) - (forget _) }, - { introI, - exact (infer_instance : (preadditive_yoneda_obj J ⋙ forget _).preserves_epimorphisms) } -end - -end preadditive - section enough_injectives variable [enough_injectives C] diff --git a/src/category_theory/preadditive/opposite.lean b/src/category_theory/preadditive/opposite.lean index 045521e03a507..1dd70b5fc7952 100644 --- a/src/category_theory/preadditive/opposite.lean +++ b/src/category_theory/preadditive/opposite.lean @@ -3,7 +3,6 @@ Copyright (c) 2021 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison, Adam Topaz, Johan Commelin, Joël Riou -/ -import category_theory.preadditive.basic import category_theory.preadditive.additive_functor import logic.equiv.transfer_instance diff --git a/src/category_theory/preadditive/projective.lean b/src/category_theory/preadditive/projective.lean index c4e4dfefe08d7..25bd3d923a45b 100644 --- a/src/category_theory/preadditive/projective.lean +++ b/src/category_theory/preadditive/projective.lean @@ -4,12 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Markus Himmel, Scott Morrison -/ import algebra.homology.exact -import category_theory.types import category_theory.limits.shapes.biproducts import category_theory.adjunction.limits -import category_theory.preadditive.yoneda.basic -import algebra.category.Group.epi_mono -import algebra.category.Module.epi_mono +import category_theory.limits.preserves.finite /-! # Projective objects and categories with enough projectives @@ -140,33 +137,6 @@ lemma projective_iff_preserves_epimorphisms_coyoneda_obj (P : C) : λ h, ⟨λ E X f e he, by exactI (epi_iff_surjective _).1 (infer_instance : epi ((coyoneda.obj (op P)).map e)) f⟩⟩ -section preadditive -variables [preadditive C] - -lemma projective_iff_preserves_epimorphisms_preadditive_coyoneda_obj (P : C) : - projective P ↔ (preadditive_coyoneda.obj (op P)).preserves_epimorphisms := -begin - rw projective_iff_preserves_epimorphisms_coyoneda_obj, - refine ⟨λ (h : (preadditive_coyoneda.obj (op P) ⋙ (forget _)).preserves_epimorphisms), _, _⟩, - { exactI functor.preserves_epimorphisms_of_preserves_of_reflects (preadditive_coyoneda.obj (op P)) - (forget _) }, - { introI, - exact (infer_instance : (preadditive_coyoneda.obj (op P) ⋙ forget _).preserves_epimorphisms) } -end - -lemma projective_iff_preserves_epimorphisms_preadditive_coyoneda_obj' (P : C) : - projective P ↔ (preadditive_coyoneda_obj (op P)).preserves_epimorphisms := -begin - rw projective_iff_preserves_epimorphisms_coyoneda_obj, - refine ⟨λ (h : (preadditive_coyoneda_obj (op P) ⋙ (forget _)).preserves_epimorphisms), _, _⟩, - { exactI functor.preserves_epimorphisms_of_preserves_of_reflects (preadditive_coyoneda_obj (op P)) - (forget _) }, - { introI, - exact (infer_instance : (preadditive_coyoneda_obj (op P) ⋙ forget _).preserves_epimorphisms) } -end - -end preadditive - section enough_projectives variables [enough_projectives C] diff --git a/src/category_theory/preadditive/yoneda/injective.lean b/src/category_theory/preadditive/yoneda/injective.lean new file mode 100644 index 0000000000000..d26a0c76459f4 --- /dev/null +++ b/src/category_theory/preadditive/yoneda/injective.lean @@ -0,0 +1,53 @@ +/- +Copyright (c) 2020 Markus Himmel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Markus Himmel, Scott Morrison +-/ +import category_theory.preadditive.yoneda.basic +import category_theory.preadditive.injective +import algebra.category.Group.epi_mono +import algebra.category.Module.epi_mono + +/-! +An object is injective iff the preadditive yoneda functor on it preserves epimorphisms. +-/ + +universes v u + +open opposite + +namespace category_theory +variables {C : Type u} [category.{v} C] + +section preadditive +variables [preadditive C] + +namespace injective + +lemma injective_iff_preserves_epimorphisms_preadditive_yoneda_obj (J : C) : + injective J ↔ (preadditive_yoneda.obj J).preserves_epimorphisms := +begin + rw injective_iff_preserves_epimorphisms_yoneda_obj, + refine ⟨λ (h : (preadditive_yoneda.obj J ⋙ (forget _)).preserves_epimorphisms), _, _⟩, + { exactI functor.preserves_epimorphisms_of_preserves_of_reflects (preadditive_yoneda.obj J) + (forget _) }, + { introI, + exact (infer_instance : (preadditive_yoneda.obj J ⋙ forget _).preserves_epimorphisms) } +end + +lemma injective_iff_preserves_epimorphisms_preadditive_yoneda_obj' (J : C) : + injective J ↔ (preadditive_yoneda_obj J).preserves_epimorphisms := +begin + rw injective_iff_preserves_epimorphisms_yoneda_obj, + refine ⟨λ (h : (preadditive_yoneda_obj J ⋙ (forget _)).preserves_epimorphisms), _, _⟩, + { exactI functor.preserves_epimorphisms_of_preserves_of_reflects (preadditive_yoneda_obj J) + (forget _) }, + { introI, + exact (infer_instance : (preadditive_yoneda_obj J ⋙ forget _).preserves_epimorphisms) } +end + +end injective + +end preadditive + +end category_theory diff --git a/src/category_theory/preadditive/yoneda/projective.lean b/src/category_theory/preadditive/yoneda/projective.lean new file mode 100644 index 0000000000000..886bb6d2a5a87 --- /dev/null +++ b/src/category_theory/preadditive/yoneda/projective.lean @@ -0,0 +1,53 @@ +/- +Copyright (c) 2020 Markus Himmel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Markus Himmel, Scott Morrison +-/ +import category_theory.preadditive.yoneda.basic +import category_theory.preadditive.projective +import algebra.category.Group.epi_mono +import algebra.category.Module.epi_mono + +/-! +An object is projective iff the preadditive coyoneda functor on it preserves epimorphisms. +-/ + +universes v u + +open opposite + +namespace category_theory +variables {C : Type u} [category.{v} C] + +section preadditive +variables [preadditive C] + +namespace projective + +lemma projective_iff_preserves_epimorphisms_preadditive_coyoneda_obj (P : C) : + projective P ↔ (preadditive_coyoneda.obj (op P)).preserves_epimorphisms := +begin + rw projective_iff_preserves_epimorphisms_coyoneda_obj, + refine ⟨λ (h : (preadditive_coyoneda.obj (op P) ⋙ (forget _)).preserves_epimorphisms), _, _⟩, + { exactI functor.preserves_epimorphisms_of_preserves_of_reflects (preadditive_coyoneda.obj (op P)) + (forget _) }, + { introI, + exact (infer_instance : (preadditive_coyoneda.obj (op P) ⋙ forget _).preserves_epimorphisms) } +end + +lemma projective_iff_preserves_epimorphisms_preadditive_coyoneda_obj' (P : C) : + projective P ↔ (preadditive_coyoneda_obj (op P)).preserves_epimorphisms := +begin + rw projective_iff_preserves_epimorphisms_coyoneda_obj, + refine ⟨λ (h : (preadditive_coyoneda_obj (op P) ⋙ (forget _)).preserves_epimorphisms), _, _⟩, + { exactI functor.preserves_epimorphisms_of_preserves_of_reflects (preadditive_coyoneda_obj (op P)) + (forget _) }, + { introI, + exact (infer_instance : (preadditive_coyoneda_obj (op P) ⋙ forget _).preserves_epimorphisms) } +end + +end projective + +end preadditive + +end category_theory From 94d4e70e97c36c896cb70fb42821acfed040de60 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Thu, 30 Mar 2023 17:06:30 +0000 Subject: [PATCH 0738/1291] chore(category_theory): removed stupid example (#18701) --- src/category_theory/triangulated/rotate.lean | 2 -- 1 file changed, 2 deletions(-) diff --git a/src/category_theory/triangulated/rotate.lean b/src/category_theory/triangulated/rotate.lean index ebc422b3e55e8..b4e212b14f6b9 100644 --- a/src/category_theory/triangulated/rotate.lean +++ b/src/category_theory/triangulated/rotate.lean @@ -86,8 +86,6 @@ def rotate : triangle C ⥤ triangle C := hom₃ := f.hom₁⟦1⟧', comm₃' := by { dsimp, simp only [comp_neg, neg_comp, ← functor.map_comp, f.comm₁], }, }, } -example : ℕ:= 42 - /-- The inverse rotation of triangles gives an endofunctor on the category of triangles in `C`. -/ From 6e272cd89fa32c72a25dbefd319394c48dce1576 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 30 Mar 2023 17:06:31 +0000 Subject: [PATCH 0739/1291] refactor(analysis/inner_product_space): split `submodule.orthogonal` to a new file (#18706) This file is pretty long, and this seems to split out naturally. The lemmas are copied with minimal modification; the only change is that everything is now wrapped in `namespace submodule` rather than having `submodule.` prefixing every lemma. --- src/analysis/inner_product_space/basic.lean | 178 --------------- .../inner_product_space/orthogonal.lean | 202 ++++++++++++++++++ .../inner_product_space/projection.lean | 1 + 3 files changed, 203 insertions(+), 178 deletions(-) create mode 100644 src/analysis/inner_product_space/orthogonal.lean diff --git a/src/analysis/inner_product_space/basic.lean b/src/analysis/inner_product_space/basic.lean index 4a82db4b3094e..557e331d90c23 100644 --- a/src/analysis/inner_product_space/basic.lean +++ b/src/analysis/inner_product_space/basic.lean @@ -39,9 +39,6 @@ product structure on `n → 𝕜` for `𝕜 = ℝ` or `ℂ`, see `euclidean_spac the sum of the norm-squares of the inner products `⟪v i, x⟫` is no more than the norm-square of `x`. For the existence of orthonormal bases, Hilbert bases, etc., see the file `analysis.inner_product_space.projection`. -- The `orthogonal_complement` of a submodule `K` is defined, and basic API established. Some of - the more subtle results about the orthogonal complement are delayed to - `analysis.inner_product_space.projection`. ## Notation @@ -49,8 +46,6 @@ We globally denote the real and complex inner products by `⟪·, ·⟫_ℝ` and We also provide two notation namespaces: `real_inner_product_space`, `complex_inner_product_space`, which respectively introduce the plain notation `⟪·, ·⟫` for the real and complex inner product. -The orthogonal complement of a submodule `K` is denoted by `Kᗮ`. - ## Implementation notes We choose the convention that inner products are conjugate linear in the first argument and linear @@ -2216,179 +2211,6 @@ by simp only [continuous_linear_map.map_smul, continuous_linear_map.re_apply_inn end re_apply_inner_self -/-! ### The orthogonal complement -/ - -section orthogonal -variables (K : submodule 𝕜 E) - -/-- The subspace of vectors orthogonal to a given subspace. -/ -def submodule.orthogonal : submodule 𝕜 E := -{ carrier := {v | ∀ u ∈ K, ⟪u, v⟫ = 0}, - zero_mem' := λ _ _, inner_zero_right _, - add_mem' := λ x y hx hy u hu, by rw [inner_add_right, hx u hu, hy u hu, add_zero], - smul_mem' := λ c x hx u hu, by rw [inner_smul_right, hx u hu, mul_zero] } - -notation K`ᗮ`:1200 := submodule.orthogonal K - -/-- When a vector is in `Kᗮ`. -/ -lemma submodule.mem_orthogonal (v : E) : v ∈ Kᗮ ↔ ∀ u ∈ K, ⟪u, v⟫ = 0 := iff.rfl - -/-- When a vector is in `Kᗮ`, with the inner product the -other way round. -/ -lemma submodule.mem_orthogonal' (v : E) : v ∈ Kᗮ ↔ ∀ u ∈ K, ⟪v, u⟫ = 0 := -by simp_rw [submodule.mem_orthogonal, inner_eq_zero_symm] - -variables {K} - -/-- A vector in `K` is orthogonal to one in `Kᗮ`. -/ -lemma submodule.inner_right_of_mem_orthogonal {u v : E} (hu : u ∈ K) (hv : v ∈ Kᗮ) : ⟪u, v⟫ = 0 := -(K.mem_orthogonal v).1 hv u hu - -/-- A vector in `Kᗮ` is orthogonal to one in `K`. -/ -lemma submodule.inner_left_of_mem_orthogonal {u v : E} (hu : u ∈ K) (hv : v ∈ Kᗮ) : ⟪v, u⟫ = 0 := -by rw [inner_eq_zero_symm]; exact submodule.inner_right_of_mem_orthogonal hu hv - -/-- A vector is in `(𝕜 ∙ u)ᗮ` iff it is orthogonal to `u`. -/ -lemma submodule.mem_orthogonal_singleton_iff_inner_right {u v : E} : v ∈ (𝕜 ∙ u)ᗮ ↔ ⟪u, v⟫ = 0 := -begin - refine ⟨submodule.inner_right_of_mem_orthogonal (submodule.mem_span_singleton_self u), _⟩, - intros hv w hw, - rw submodule.mem_span_singleton at hw, - obtain ⟨c, rfl⟩ := hw, - simp [inner_smul_left, hv], -end - -/-- A vector in `(𝕜 ∙ u)ᗮ` is orthogonal to `u`. -/ -lemma submodule.mem_orthogonal_singleton_iff_inner_left {u v : E} : v ∈ (𝕜 ∙ u)ᗮ ↔ ⟪v, u⟫ = 0 := -by rw [submodule.mem_orthogonal_singleton_iff_inner_right, inner_eq_zero_symm] - -lemma submodule.sub_mem_orthogonal_of_inner_left {x y : E} - (h : ∀ (v : K), ⟪x, v⟫ = ⟪y, v⟫) : x - y ∈ Kᗮ := -begin - rw submodule.mem_orthogonal', - intros u hu, - rw [inner_sub_left, sub_eq_zero], - exact h ⟨u, hu⟩, -end - -lemma submodule.sub_mem_orthogonal_of_inner_right {x y : E} - (h : ∀ (v : K), ⟪(v : E), x⟫ = ⟪(v : E), y⟫) : x - y ∈ Kᗮ := -begin - intros u hu, - rw [inner_sub_right, sub_eq_zero], - exact h ⟨u, hu⟩, -end - -variables (K) - -/-- `K` and `Kᗮ` have trivial intersection. -/ -lemma submodule.inf_orthogonal_eq_bot : K ⊓ Kᗮ = ⊥ := -begin - rw submodule.eq_bot_iff, - intros x, - rw submodule.mem_inf, - exact λ ⟨hx, ho⟩, inner_self_eq_zero.1 (ho x hx) -end - -/-- `K` and `Kᗮ` have trivial intersection. -/ -lemma submodule.orthogonal_disjoint : disjoint K Kᗮ := -by simp [disjoint_iff, K.inf_orthogonal_eq_bot] - -/-- `Kᗮ` can be characterized as the intersection of the kernels of the operations of -inner product with each of the elements of `K`. -/ -lemma orthogonal_eq_inter : Kᗮ = ⨅ v : K, linear_map.ker (innerSL 𝕜 (v : E)) := -begin - apply le_antisymm, - { rw le_infi_iff, - rintros ⟨v, hv⟩ w hw, - simpa using hw _ hv }, - { intros v hv w hw, - simp only [submodule.mem_infi] at hv, - exact hv ⟨w, hw⟩ } -end - -/-- The orthogonal complement of any submodule `K` is closed. -/ -lemma submodule.is_closed_orthogonal : is_closed (Kᗮ : set E) := -begin - rw orthogonal_eq_inter K, - have := λ v : K, continuous_linear_map.is_closed_ker (innerSL 𝕜 (v : E)), - convert is_closed_Inter this, - simp only [submodule.infi_coe], -end - -/-- In a complete space, the orthogonal complement of any submodule `K` is complete. -/ -instance [complete_space E] : complete_space Kᗮ := K.is_closed_orthogonal.complete_space_coe - -variables (𝕜 E) - -/-- `submodule.orthogonal` gives a `galois_connection` between -`submodule 𝕜 E` and its `order_dual`. -/ -lemma submodule.orthogonal_gc : - @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)⟩ - -variables {𝕜 E} - -/-- `submodule.orthogonal` reverses the `≤` ordering of two -subspaces. -/ -lemma submodule.orthogonal_le {K₁ K₂ : submodule 𝕜 E} (h : K₁ ≤ K₂) : K₂ᗮ ≤ K₁ᗮ := -(submodule.orthogonal_gc 𝕜 E).monotone_l h - -/-- `submodule.orthogonal.orthogonal` preserves the `≤` ordering of two -subspaces. -/ -lemma submodule.orthogonal_orthogonal_monotone {K₁ K₂ : submodule 𝕜 E} (h : K₁ ≤ K₂) : - K₁ᗮᗮ ≤ K₂ᗮᗮ := -submodule.orthogonal_le (submodule.orthogonal_le h) - -/-- `K` is contained in `Kᗮᗮ`. -/ -lemma submodule.le_orthogonal_orthogonal : K ≤ Kᗮᗮ := (submodule.orthogonal_gc 𝕜 E).le_u_l _ - -/-- The inf of two orthogonal subspaces equals the subspace orthogonal -to the sup. -/ -lemma submodule.inf_orthogonal (K₁ K₂ : submodule 𝕜 E) : K₁ᗮ ⊓ K₂ᗮ = (K₁ ⊔ K₂)ᗮ := -(submodule.orthogonal_gc 𝕜 E).l_sup.symm - -/-- The inf of an indexed family of orthogonal subspaces equals the -subspace orthogonal to the sup. -/ -lemma submodule.infi_orthogonal {ι : Type*} (K : ι → submodule 𝕜 E) : (⨅ i, (K i)ᗮ) = (supr K)ᗮ := -(submodule.orthogonal_gc 𝕜 E).l_supr.symm - -/-- The inf of a set of orthogonal subspaces equals the subspace orthogonal to the sup. -/ -lemma submodule.Inf_orthogonal (s : set $ submodule 𝕜 E) : (⨅ K ∈ s, Kᗮ) = (Sup s)ᗮ := -(submodule.orthogonal_gc 𝕜 E).l_Sup.symm - -@[simp] lemma submodule.top_orthogonal_eq_bot : (⊤ : submodule 𝕜 E)ᗮ = ⊥ := -begin - ext, - rw [submodule.mem_bot, submodule.mem_orthogonal], - exact ⟨λ h, inner_self_eq_zero.mp (h x submodule.mem_top), by { rintro rfl, simp }⟩ -end - -@[simp] lemma submodule.bot_orthogonal_eq_top : (⊥ : submodule 𝕜 E)ᗮ = ⊤ := -begin - rw [← submodule.top_orthogonal_eq_bot, eq_top_iff], - exact submodule.le_orthogonal_orthogonal ⊤ -end - -@[simp] lemma submodule.orthogonal_eq_top_iff : Kᗮ = ⊤ ↔ K = ⊥ := -begin - refine ⟨_, by { rintro rfl, exact submodule.bot_orthogonal_eq_top }⟩, - intro h, - have : K ⊓ Kᗮ = ⊥ := K.orthogonal_disjoint.eq_bot, - rwa [h, inf_comm, top_inf_eq] at this -end - -lemma submodule.orthogonal_family_self : - orthogonal_family 𝕜 (λ b, ↥(cond b K Kᗮ)) (λ b, (cond b K Kᗮ).subtypeₗᵢ) -| tt tt := absurd rfl -| tt ff := λ _ x y, submodule.inner_right_of_mem_orthogonal x.prop y.prop -| ff tt := λ _ x y, submodule.inner_left_of_mem_orthogonal y.prop x.prop -| ff ff := absurd rfl - -end orthogonal - namespace uniform_space.completion open uniform_space function diff --git a/src/analysis/inner_product_space/orthogonal.lean b/src/analysis/inner_product_space/orthogonal.lean new file mode 100644 index 0000000000000..e74f2e7011cab --- /dev/null +++ b/src/analysis/inner_product_space/orthogonal.lean @@ -0,0 +1,202 @@ +/- +Copyright (c) 2019 Zhouhang Zhou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Zhouhang Zhou, Sébastien Gouëzel, Frédéric Dupuis +-/ +import algebra.direct_sum.module +import analysis.complex.basic +import analysis.convex.uniform +import analysis.normed_space.completion +import analysis.normed_space.bounded_linear_maps +import linear_algebra.bilinear_form + +import analysis.inner_product_space.basic + +/-! +# Orthogonal complements of submodules + +In this file, the `orthogonal` complement of a submodule `K` is defined, and basic API established. +Some of the more subtle results about the orthogonal complement are delayed to +`analysis.inner_product_space.projection`. + +## Notation + +The orthogonal complement of a submodule `K` is denoted by `Kᗮ`. +-/ + +variables {𝕜 E F : Type*} [is_R_or_C 𝕜] +variables [normed_add_comm_group E] [inner_product_space 𝕜 E] +variables [normed_add_comm_group F] [inner_product_space ℝ F] +local notation `⟪`x`, `y`⟫` := @inner 𝕜 _ _ x y + +namespace submodule + +variables (K : submodule 𝕜 E) + +/-- The subspace of vectors orthogonal to a given subspace. -/ +def orthogonal : submodule 𝕜 E := +{ carrier := {v | ∀ u ∈ K, ⟪u, v⟫ = 0}, + zero_mem' := λ _ _, inner_zero_right _, + add_mem' := λ x y hx hy u hu, by rw [inner_add_right, hx u hu, hy u hu, add_zero], + smul_mem' := λ c x hx u hu, by rw [inner_smul_right, hx u hu, mul_zero] } + +notation K`ᗮ`:1200 := orthogonal K + +/-- When a vector is in `Kᗮ`. -/ +lemma mem_orthogonal (v : E) : v ∈ Kᗮ ↔ ∀ u ∈ K, ⟪u, v⟫ = 0 := iff.rfl + +/-- When a vector is in `Kᗮ`, with the inner product the +other way round. -/ +lemma mem_orthogonal' (v : E) : v ∈ Kᗮ ↔ ∀ u ∈ K, ⟪v, u⟫ = 0 := +by simp_rw [mem_orthogonal, inner_eq_zero_symm] + +variables {K} + +/-- A vector in `K` is orthogonal to one in `Kᗮ`. -/ +lemma inner_right_of_mem_orthogonal {u v : E} (hu : u ∈ K) (hv : v ∈ Kᗮ) : ⟪u, v⟫ = 0 := +(K.mem_orthogonal v).1 hv u hu + +/-- A vector in `Kᗮ` is orthogonal to one in `K`. -/ +lemma inner_left_of_mem_orthogonal {u v : E} (hu : u ∈ K) (hv : v ∈ Kᗮ) : ⟪v, u⟫ = 0 := +by rw [inner_eq_zero_symm]; exact inner_right_of_mem_orthogonal hu hv + +/-- A vector is in `(𝕜 ∙ u)ᗮ` iff it is orthogonal to `u`. -/ +lemma mem_orthogonal_singleton_iff_inner_right {u v : E} : v ∈ (𝕜 ∙ u)ᗮ ↔ ⟪u, v⟫ = 0 := +begin + refine ⟨inner_right_of_mem_orthogonal (mem_span_singleton_self u), _⟩, + intros hv w hw, + rw mem_span_singleton at hw, + obtain ⟨c, rfl⟩ := hw, + simp [inner_smul_left, hv], +end + +/-- A vector in `(𝕜 ∙ u)ᗮ` is orthogonal to `u`. -/ +lemma mem_orthogonal_singleton_iff_inner_left {u v : E} : v ∈ (𝕜 ∙ u)ᗮ ↔ ⟪v, u⟫ = 0 := +by rw [mem_orthogonal_singleton_iff_inner_right, inner_eq_zero_symm] + +lemma sub_mem_orthogonal_of_inner_left {x y : E} + (h : ∀ (v : K), ⟪x, v⟫ = ⟪y, v⟫) : x - y ∈ Kᗮ := +begin + rw mem_orthogonal', + intros u hu, + rw [inner_sub_left, sub_eq_zero], + exact h ⟨u, hu⟩, +end + +lemma sub_mem_orthogonal_of_inner_right {x y : E} + (h : ∀ (v : K), ⟪(v : E), x⟫ = ⟪(v : E), y⟫) : x - y ∈ Kᗮ := +begin + intros u hu, + rw [inner_sub_right, sub_eq_zero], + exact h ⟨u, hu⟩, +end + +variables (K) + +/-- `K` and `Kᗮ` have trivial intersection. -/ +lemma inf_orthogonal_eq_bot : K ⊓ Kᗮ = ⊥ := +begin + rw eq_bot_iff, + intros x, + rw mem_inf, + exact λ ⟨hx, ho⟩, inner_self_eq_zero.1 (ho x hx) +end + +/-- `K` and `Kᗮ` have trivial intersection. -/ +lemma orthogonal_disjoint : disjoint K Kᗮ := +by simp [disjoint_iff, K.inf_orthogonal_eq_bot] + +/-- `Kᗮ` can be characterized as the intersection of the kernels of the operations of +inner product with each of the elements of `K`. -/ +lemma orthogonal_eq_inter : Kᗮ = ⨅ v : K, linear_map.ker (innerSL 𝕜 (v : E)) := +begin + apply le_antisymm, + { rw le_infi_iff, + rintros ⟨v, hv⟩ w hw, + simpa using hw _ hv }, + { intros v hv w hw, + simp only [mem_infi] at hv, + exact hv ⟨w, hw⟩ } +end + +/-- The orthogonal complement of any submodule `K` is closed. -/ +lemma is_closed_orthogonal : is_closed (Kᗮ : set E) := +begin + rw orthogonal_eq_inter K, + have := λ v : K, continuous_linear_map.is_closed_ker (innerSL 𝕜 (v : E)), + convert is_closed_Inter this, + simp only [infi_coe], +end + +/-- In a complete space, the orthogonal complement of any submodule `K` is complete. -/ +instance [complete_space E] : complete_space Kᗮ := K.is_closed_orthogonal.complete_space_coe + +variables (𝕜 E) + +/-- `orthogonal` gives a `galois_connection` between +`submodule 𝕜 E` and its `order_dual`. -/ +lemma orthogonal_gc : + @galois_connection (submodule 𝕜 E) (submodule 𝕜 E)ᵒᵈ _ _ + orthogonal orthogonal := +λ K₁ K₂, ⟨λ h v hv u hu, inner_left_of_mem_orthogonal hv (h hu), + λ h v hv u hu, inner_left_of_mem_orthogonal hv (h hu)⟩ + +variables {𝕜 E} + +/-- `orthogonal` reverses the `≤` ordering of two +subspaces. -/ +lemma orthogonal_le {K₁ K₂ : submodule 𝕜 E} (h : K₁ ≤ K₂) : K₂ᗮ ≤ K₁ᗮ := +(orthogonal_gc 𝕜 E).monotone_l h + +/-- `orthogonal.orthogonal` preserves the `≤` ordering of two +subspaces. -/ +lemma orthogonal_orthogonal_monotone {K₁ K₂ : submodule 𝕜 E} (h : K₁ ≤ K₂) : + K₁ᗮᗮ ≤ K₂ᗮᗮ := +orthogonal_le (orthogonal_le h) + +/-- `K` is contained in `Kᗮᗮ`. -/ +lemma le_orthogonal_orthogonal : K ≤ Kᗮᗮ := (orthogonal_gc 𝕜 E).le_u_l _ + +/-- The inf of two orthogonal subspaces equals the subspace orthogonal +to the sup. -/ +lemma inf_orthogonal (K₁ K₂ : submodule 𝕜 E) : K₁ᗮ ⊓ K₂ᗮ = (K₁ ⊔ K₂)ᗮ := +(orthogonal_gc 𝕜 E).l_sup.symm + +/-- The inf of an indexed family of orthogonal subspaces equals the +subspace orthogonal to the sup. -/ +lemma infi_orthogonal {ι : Type*} (K : ι → submodule 𝕜 E) : (⨅ i, (K i)ᗮ) = (supr K)ᗮ := +(orthogonal_gc 𝕜 E).l_supr.symm + +/-- The inf of a set of orthogonal subspaces equals the subspace orthogonal to the sup. -/ +lemma Inf_orthogonal (s : set $ submodule 𝕜 E) : (⨅ K ∈ s, Kᗮ) = (Sup s)ᗮ := +(orthogonal_gc 𝕜 E).l_Sup.symm + +@[simp] lemma top_orthogonal_eq_bot : (⊤ : submodule 𝕜 E)ᗮ = ⊥ := +begin + ext, + rw [mem_bot, mem_orthogonal], + exact ⟨λ h, inner_self_eq_zero.mp (h x mem_top), by { rintro rfl, simp }⟩ +end + +@[simp] lemma bot_orthogonal_eq_top : (⊥ : submodule 𝕜 E)ᗮ = ⊤ := +begin + rw [← top_orthogonal_eq_bot, eq_top_iff], + exact le_orthogonal_orthogonal ⊤ +end + +@[simp] lemma orthogonal_eq_top_iff : Kᗮ = ⊤ ↔ K = ⊥ := +begin + refine ⟨_, by { rintro rfl, exact bot_orthogonal_eq_top }⟩, + intro h, + have : K ⊓ Kᗮ = ⊥ := K.orthogonal_disjoint.eq_bot, + rwa [h, inf_comm, top_inf_eq] at this +end + +lemma orthogonal_family_self : + orthogonal_family 𝕜 (λ b, ↥(cond b K Kᗮ)) (λ b, (cond b K Kᗮ).subtypeₗᵢ) +| tt tt := absurd rfl +| tt ff := λ _ x y, inner_right_of_mem_orthogonal x.prop y.prop +| ff tt := λ _ x y, inner_left_of_mem_orthogonal y.prop x.prop +| ff ff := absurd rfl + +end submodule diff --git a/src/analysis/inner_product_space/projection.lean b/src/analysis/inner_product_space/projection.lean index e2e124bec373a..04da7b4e7a6b1 100644 --- a/src/analysis/inner_product_space/projection.lean +++ b/src/analysis/inner_product_space/projection.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Zhouhang Zhou, Frédéric Dupuis, Heather Macbeth -/ import analysis.convex.basic +import analysis.inner_product_space.orthogonal import analysis.inner_product_space.symmetric import analysis.normed_space.is_R_or_C import data.is_R_or_C.lemmas From ce11c3c2a285bbe6937e26d9792fda4e51f3fe1a Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 30 Mar 2023 18:52:18 +0000 Subject: [PATCH 0740/1291] refactor(linear_algebra/{multilinear,pi_tensor_product}): remove the decidable_eq argument (#10140) There is no need to include this argument in the type of `multilinear_map`, as it is only used to state the propositions. Instead, we move it into a binder in the `map_add` and `map_smul` fields. The same trick is done for the relation for `pi_tensor_product`. The benefit here is pretty marginal; the main one is that `mutlilinear_map.mk_pi_algebra` no longer requires a `decidable_eq I` instance. However, it still requires `fintype I`, and in practice all computably finite types are also computably decidable. This does at least mean that `0 : multilinear_map ...` rightfully no longer requires decidable equality! The downside of this PR is that the `map_add'` and `map_smul'` fields are often more annoying to prove as there are sometimes duplicate decidable instances to eliminate. [Zulip thread](https://leanprover.zulipchat.com/#narrow/stream/113488-general/topic/A.20possible.20diamond/near/260179946) --- src/analysis/analytic/composition.lean | 10 +- src/analysis/analytic/inverse.lean | 2 +- .../normed_space/bounded_linear_maps.lean | 6 +- src/analysis/normed_space/multilinear.lean | 43 ++-- src/linear_algebra/alternating.lean | 50 ++-- src/linear_algebra/determinant.lean | 21 +- .../exterior_algebra/of_alternating.lean | 2 +- src/linear_algebra/multilinear/basic.lean | 238 +++++++++++------- src/linear_algebra/multilinear/basis.lean | 2 +- .../multilinear/finite_dimensional.lean | 2 +- .../multilinear/tensor_product.lean | 13 +- src/linear_algebra/orientation.lean | 21 +- src/linear_algebra/pi_tensor_product.lean | 44 ++-- src/linear_algebra/tensor_power.lean | 3 +- src/topology/algebra/module/multilinear.lean | 22 +- 15 files changed, 281 insertions(+), 198 deletions(-) diff --git a/src/analysis/analytic/composition.lean b/src/analysis/analytic/composition.lean index 37f8e87141756..72c45ccf16d31 100644 --- a/src/analysis/analytic/composition.lean +++ b/src/analysis/analytic/composition.lean @@ -181,10 +181,12 @@ def comp_along_composition {n : ℕ} (f : continuous_multilinear_map 𝕜 (λ (i : fin c.length), F) G) : continuous_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], + map_add' := λ _ v i x y, by + { cases subsingleton.elim ‹_› (fin.decidable_eq _), + simp only [apply_composition_update, continuous_multilinear_map.map_add] }, + map_smul' := λ _ v i c x, by + { cases subsingleton.elim ‹_› (fin.decidable_eq _), + 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 _, } diff --git a/src/analysis/analytic/inverse.lean b/src/analysis/analytic/inverse.lean index f3c4ca2afb134..82fb6aaad1073 100644 --- a/src/analysis/analytic/inverse.lean +++ b/src/analysis/analytic/inverse.lean @@ -393,7 +393,7 @@ begin (λ (k : ℕ), (fintype.pi_finset (λ (i : fin k), Ico 1 n) : finset (fin k → ℕ))) (λ n e, ∏ (j : fin n), r * (a ^ e j * p (e j)))], apply sum_congr rfl (λ j hj, _), - simp only [← @multilinear_map.mk_pi_algebra_apply ℝ (fin j) _ _ ℝ], + simp only [← @multilinear_map.mk_pi_algebra_apply ℝ (fin j) _ ℝ], simp only [← multilinear_map.map_sum_finset (multilinear_map.mk_pi_algebra ℝ (fin j) ℝ) (λ k (m : ℕ), r * (a ^ m * p m))], simp only [multilinear_map.mk_pi_algebra_apply], diff --git a/src/analysis/normed_space/bounded_linear_maps.lean b/src/analysis/normed_space/bounded_linear_maps.lean index 150fb5a3c6a41..d7f734a30866d 100644 --- a/src/analysis/normed_space/bounded_linear_maps.lean +++ b/src/analysis/normed_space/bounded_linear_maps.lean @@ -49,7 +49,7 @@ artifact, really. -/ noncomputable theory -open_locale classical big_operators topology +open_locale big_operators topology open filter (tendsto) metric continuous_linear_map @@ -183,7 +183,7 @@ end end is_bounded_linear_map section -variables {ι : Type*} [decidable_eq ι] [fintype ι] +variables {ι : Type*} [fintype ι] /-- Taking the cartesian product of two continuous multilinear maps is a bounded linear operation. -/ @@ -441,7 +441,7 @@ lemma is_bounded_bilinear_map_smul_right : /-- The composition of a continuous linear map with a continuous multilinear map is a bounded bilinear operation. -/ lemma is_bounded_bilinear_map_comp_multilinear {ι : Type*} {E : ι → Type*} -[decidable_eq ι] [fintype ι] [∀ i, normed_add_comm_group (E i)] [∀ i, normed_space 𝕜 (E i)] : + [fintype ι] [∀ i, normed_add_comm_group (E i)] [∀ i, normed_space 𝕜 (E i)] : is_bounded_bilinear_map 𝕜 (λ p : (F →L[𝕜] G) × (continuous_multilinear_map 𝕜 E F), p.1.comp_continuous_multilinear_map p.2) := (comp_continuous_multilinear_mapL 𝕜 E F G).is_bounded_bilinear_map diff --git a/src/analysis/normed_space/multilinear.lean b/src/analysis/normed_space/multilinear.lean index fab1ad5232064..86a981e9cc103 100644 --- a/src/analysis/normed_space/multilinear.lean +++ b/src/analysis/normed_space/multilinear.lean @@ -54,7 +54,7 @@ approach, it turns out that direct proofs are easier and more efficient. -/ noncomputable theory -open_locale classical big_operators nnreal +open_locale big_operators nnreal open finset metric local attribute [instance, priority 1001] @@ -77,7 +77,7 @@ universes u v v' wE wE₁ wE' wEi wG wG' variables {𝕜 : Type u} {ι : Type v} {ι' : Type v'} {n : ℕ} {E : ι → Type wE} {E₁ : ι → Type wE₁} {E' : ι' → Type wE'} {Ei : fin n.succ → Type wEi} {G : Type wG} {G' : Type wG'} - [decidable_eq ι] [fintype ι] [decidable_eq ι'] [fintype ι'] [nontrivially_normed_field 𝕜] + [fintype ι] [fintype ι'] [nontrivially_normed_field 𝕜] [Π i, normed_add_comm_group (E i)] [Π i, normed_space 𝕜 (E i)] [Π i, normed_add_comm_group (E₁ i)] [Π i, normed_space 𝕜 (E₁ i)] [Π i, normed_add_comm_group (E' i)] [Π i, normed_space 𝕜 (E' i)] @@ -137,7 +137,7 @@ using the multilinearity. Here, we give a precise but hard to use version. See `‖f m - f m'‖ ≤ C * ‖m 1 - m' 1‖ * max ‖m 2‖ ‖m' 2‖ * max ‖m 3‖ ‖m' 3‖ * ... * max ‖m n‖ ‖m' n‖ + ...`, where the other terms in the sum are the same products where `1` is replaced by any `i`. -/ -lemma norm_image_sub_le_of_bound' {C : ℝ} (hC : 0 ≤ C) +lemma norm_image_sub_le_of_bound' [decidable_eq ι] {C : ℝ} (hC : 0 ≤ C) (H : ∀ m, ‖f m‖ ≤ C * ∏ i, ‖m i‖) (m₁ m₂ : Πi, E i) : ‖f m₁ - f m₂‖ ≤ C * ∑ i, ∏ j, if j = i then ‖m₁ i - m₂ i‖ else max ‖m₁ j‖ ‖m₂ j‖ := @@ -182,6 +182,7 @@ lemma norm_image_sub_le_of_bound {C : ℝ} (hC : 0 ≤ C) (H : ∀ m, ‖f m‖ ≤ C * ∏ i, ‖m i‖) (m₁ m₂ : Πi, E i) : ‖f m₁ - f m₂‖ ≤ C * (fintype.card ι) * (max ‖m₁‖ ‖m₂‖) ^ (fintype.card ι - 1) * ‖m₁ - m₂‖ := begin + letI := classical.dec_eq ι, have A : ∀ (i : ι), ∏ j, (if j = i then ‖m₁ i - m₂ i‖ else max ‖m₁ j‖ ‖m₂ j‖) ≤ ‖m₁ - m₂‖ * (max ‖m₁‖ ‖m₂‖) ^ (fintype.card ι - 1), { assume i, @@ -541,7 +542,7 @@ For a less precise but more usable version, see `norm_image_sub_le`. The bound r `‖f m - f m'‖ ≤ ‖f‖ * ‖m 1 - m' 1‖ * max ‖m 2‖ ‖m' 2‖ * max ‖m 3‖ ‖m' 3‖ * ... * max ‖m n‖ ‖m' n‖ + ...`, where the other terms in the sum are the same products where `1` is replaced by any `i`.-/ -lemma norm_image_sub_le' (m₁ m₂ : Πi, E i) : +lemma norm_image_sub_le' [decidable_eq ι] (m₁ m₂ : Πi, E i) : ‖f m₁ - f m₂‖ ≤ ‖f‖ * ∑ i, ∏ j, if j = i then ‖m₁ i - m₂ i‖ else max ‖m₁ j‖ ‖m₂ j‖ := f.to_multilinear_map.norm_image_sub_le_of_bound' (norm_nonneg _) f.le_op_norm _ _ @@ -633,13 +634,15 @@ begin -- Next, we show that this `F` is multilinear, let Fmult : multilinear_map 𝕜 E G := { to_fun := F, - map_add' := λ v i x y, begin + map_add' := λ _ v i x y, begin + resetI, have A := hF (function.update v i (x + y)), have B := (hF (function.update v i x)).add (hF (function.update v i y)), simp at A B, exact tendsto_nhds_unique A B end, - map_smul' := λ v i c x, begin + map_smul' := λ _ v i c x, begin + resetI, have A := hF (function.update v i (c • x)), have B := filter.tendsto.smul (@tendsto_const_nhds _ ℕ _ c _) (hF (function.update v i x)), simp at A B, @@ -722,7 +725,7 @@ variables {𝕜 ι} {A : Type*} [normed_comm_ring A] [normed_algebra 𝕜 A] lemma norm_mk_pi_algebra_le [nonempty ι] : ‖continuous_multilinear_map.mk_pi_algebra 𝕜 ι A‖ ≤ 1 := begin - have := λ f, @op_norm_le_bound 𝕜 ι (λ i, A) A _ _ _ _ _ _ _ f _ zero_le_one, + have := λ f, @op_norm_le_bound 𝕜 ι (λ i, A) A _ _ _ _ _ _ f _ zero_le_one, refine this _ _, intros m, simp only [continuous_multilinear_map.mk_pi_algebra_apply, one_mul], @@ -733,7 +736,7 @@ lemma norm_mk_pi_algebra_of_empty [is_empty ι] : ‖continuous_multilinear_map.mk_pi_algebra 𝕜 ι A‖ = ‖(1 : A)‖ := begin apply le_antisymm, - { have := λ f, @op_norm_le_bound 𝕜 ι (λ i, A) A _ _ _ _ _ _ _ f _ (norm_nonneg (1 : A)), + { have := λ f, @op_norm_le_bound 𝕜 ι (λ i, A) A _ _ _ _ _ _ f _ (norm_nonneg (1 : A)), refine this _ _, simp, }, { convert ratio_le_op_norm _ (λ _, (1 : A)), @@ -759,7 +762,7 @@ variables {𝕜 n} {A : Type*} [normed_ring A] [normed_algebra 𝕜 A] lemma norm_mk_pi_algebra_fin_succ_le : ‖continuous_multilinear_map.mk_pi_algebra_fin 𝕜 n.succ A‖ ≤ 1 := begin - have := λ f, @op_norm_le_bound 𝕜 (fin n.succ) (λ i, A) A _ _ _ _ _ _ _ f _ zero_le_one, + have := λ f, @op_norm_le_bound 𝕜 (fin n.succ) (λ i, A) A _ _ _ _ _ _ f _ zero_le_one, refine this _ _, intros m, simp only [continuous_multilinear_map.mk_pi_algebra_fin_apply, one_mul, list.of_fn_eq_map, @@ -781,7 +784,7 @@ lemma norm_mk_pi_algebra_fin_zero : ‖continuous_multilinear_map.mk_pi_algebra_fin 𝕜 0 A‖ = ‖(1 : A)‖ := begin refine le_antisymm _ _, - { have := λ f, @op_norm_le_bound 𝕜 (fin 0) (λ i, A) A _ _ _ _ _ _ _ f _ (norm_nonneg (1 : A)), + { have := λ f, @op_norm_le_bound 𝕜 (fin 0) (λ i, A) A _ _ _ _ _ _ f _ (norm_nonneg (1 : A)), refine this _ _, simp, }, { convert ratio_le_op_norm _ (λ _, (1 : A)), @@ -926,10 +929,10 @@ multilinear_map.mk_continuous ring_hom.id_apply] } (‖f‖ * ∏ i, ‖m i‖) $ λ x, by { rw mul_right_comm, exact (f x).le_of_op_norm_le _ (f.le_op_norm x) }, - map_add' := λ m i x y, + map_add' := λ _ m i x y, by { ext1, simp only [add_apply, continuous_multilinear_map.map_add, linear_map.coe_mk, linear_map.mk_continuous_apply]}, - map_smul' := λ m i c x, + map_smul' := λ _ m i c x, by { ext1, simp only [coe_smul', continuous_multilinear_map.map_smul, linear_map.coe_mk, linear_map.mk_continuous_apply, pi.smul_apply]} } ‖f‖ $ λ m, @@ -988,8 +991,8 @@ def mk_continuous_multilinear (f : multilinear_map 𝕜 E (multilinear_map 𝕜 continuous_multilinear_map 𝕜 E (continuous_multilinear_map 𝕜 E' G) := mk_continuous { to_fun := λ m, mk_continuous (f m) (C * ∏ i, ‖m i‖) $ H m, - map_add' := λ m i x y, by { ext1, simp }, - map_smul' := λ m i c x, by { ext1, simp } } + map_add' := λ _ m i x y, by { ext1, simp }, + map_smul' := λ _ m i c x, by { ext1, simp } } (max C 0) $ λ m, ((f m).mk_continuous_norm_le' _).trans_eq $ by { rw [max_mul_of_nonneg, zero_mul], exact prod_nonneg (λ _ _, norm_nonneg _) } @@ -1069,7 +1072,7 @@ linear_map.mk_continuous rfl lemma norm_comp_continuous_linear_mapL_le (f : Π i, E i →L[𝕜] E₁ i) : - ‖@comp_continuous_linear_mapL 𝕜 ι E E₁ G _ _ _ _ _ _ _ _ _ f‖ ≤ (∏ i, ‖f i‖) := + ‖@comp_continuous_linear_mapL 𝕜 ι E E₁ G _ _ _ _ _ _ _ _ f‖ ≤ (∏ i, ‖f i‖) := linear_map.mk_continuous_norm_le _ (prod_nonneg $ λ i _, norm_nonneg _) _ variable (G) @@ -1276,8 +1279,8 @@ def continuous_multilinear_map.uncurry_right continuous_multilinear_map 𝕜 Ei G := let f' : multilinear_map 𝕜 (λ(i : fin n), Ei i.cast_succ) (Ei (last n) →ₗ[𝕜] G) := { to_fun := λ m, (f m).to_linear_map, - map_add' := λ m i x y, by simp, - map_smul' := λ m i c x, by simp } in + map_add' := λ _ m i x y, by simp, + map_smul' := λ _ m i c x, by simp } in (@multilinear_map.uncurry_right 𝕜 n Ei G _ _ _ _ _ f').mk_continuous (‖f‖) (λm, f.norm_map_init_le m) @@ -1295,8 +1298,8 @@ def continuous_multilinear_map.curry_right let f' : multilinear_map 𝕜 (λ(i : fin n), Ei i.cast_succ) (Ei (last n) →L[𝕜] G) := { to_fun := λm, (f.to_multilinear_map.curry_right m).mk_continuous (‖f‖ * ∏ i, ‖m i‖) $ λx, f.norm_map_snoc_le m x, - map_add' := λ m i x y, by { simp, refl }, - map_smul' := λ m i c x, by { simp, refl } } in + map_add' := λ _ m i x y, by { simp, refl }, + map_smul' := λ _ m i c x, by { simp, refl } } in f'.mk_continuous (‖f‖) (λm, linear_map.mk_continuous_norm_le _ (mul_nonneg (norm_nonneg _) (prod_nonneg (λj hj, norm_nonneg _))) _) @@ -1518,8 +1521,6 @@ variables {𝕜 G G'} section -variable [decidable_eq (ι ⊕ ι')] - /-- A continuous multilinear map with variables indexed by `ι ⊕ ι'` defines a continuous multilinear map with variables indexed by `ι` taking values in the space of continuous multilinear maps with variables indexed by `ι'`. -/ diff --git a/src/linear_algebra/alternating.lean b/src/linear_algebra/alternating.lean index dbbcf8a82b4bb..73f03c6f237e9 100644 --- a/src/linear_algebra/alternating.lean +++ b/src/linear_algebra/alternating.lean @@ -53,7 +53,7 @@ variables {N : Type*} [add_comm_monoid N] [module R N] variables {M' : Type*} [add_comm_group M'] [module R M'] variables {N' : Type*} [add_comm_group N'] [module R N'] -variables {ι ι' ι'' : Type*} [decidable_eq ι] [decidable_eq ι'] [decidable_eq ι''] +variables {ι ι' ι'' : Type*} set_option old_structure_cmd true @@ -120,7 +120,8 @@ lemma coe_multilinear_map_injective : @[simp] lemma to_multilinear_map_eq_coe : f.to_multilinear_map = f := rfl @[simp] lemma coe_multilinear_map_mk (f : (ι → M) → N) (h₁ h₂ h₃) : - ((⟨f, h₁, h₂, h₃⟩ : alternating_map R M N ι) : multilinear_map R (λ i : ι, M) N) = ⟨f, h₁, h₂⟩ := + ((⟨f, h₁, h₂, h₃⟩ : alternating_map R M N ι) : multilinear_map R (λ i : ι, M) N) + = ⟨f, @h₁, @h₂⟩ := rfl end coercions @@ -130,19 +131,19 @@ end coercions These are expressed in terms of `⇑f` instead of `f.to_fun`. -/ -@[simp] lemma map_add (i : ι) (x y : M) : +@[simp] lemma map_add [decidable_eq ι] (i : ι) (x y : M) : f (update v i (x + y)) = f (update v i x) + f (update v i y) := f.to_multilinear_map.map_add' v i x y -@[simp] lemma map_sub (i : ι) (x y : M') : +@[simp] lemma map_sub [decidable_eq ι] (i : ι) (x y : M') : g' (update v' i (x - y)) = g' (update v' i x) - g' (update v' i y) := g'.to_multilinear_map.map_sub v' i x y -@[simp] lemma map_neg (i : ι) (x : M') : +@[simp] lemma map_neg [decidable_eq ι] (i : ι) (x : M') : g' (update v' i (-x)) = -g' (update v' i x) := g'.to_multilinear_map.map_neg v' i x -@[simp] lemma map_smul (i : ι) (r : R) (x : M) : +@[simp] lemma map_smul [decidable_eq ι] (i : ι) (r : R) (x : M) : f (update v i (r • x)) = r • f (update v i x) := f.to_multilinear_map.map_smul' v i r x @@ -153,7 +154,7 @@ f.map_eq_zero_of_eq' v i j h hij lemma map_coord_zero {m : ι → M} (i : ι) (h : m i = 0) : f m = 0 := f.to_multilinear_map.map_coord_zero i h -@[simp] lemma map_update_zero (m : ι → M) (i : ι) : f (update m i 0) = 0 := +@[simp] lemma map_update_zero [decidable_eq ι] (m : ι → M) (i : ι) : f (update m i 0) = 0 := f.to_multilinear_map.map_update_zero m i @[simp] lemma map_zero [nonempty ι] : f 0 = 0 := @@ -441,7 +442,7 @@ section open_locale big_operators -lemma map_update_sum {α : Type*} (t : finset α) (i : ι) (g : α → M) (m : ι → M): +lemma map_update_sum {α : Type*} [decidable_eq ι] (t : finset α) (i : ι) (g : α → M) (m : ι → M) : f (update m i (∑ a in t, g a)) = ∑ a in t, f (update m i (g a)) := f.to_multilinear_map.map_update_sum t i g m @@ -454,16 +455,16 @@ Various properties of reordered and repeated inputs which follow from `alternating_map.map_eq_zero_of_eq`. -/ -lemma map_update_self {i j : ι} (hij : i ≠ j) : +lemma map_update_self [decidable_eq ι] {i j : ι} (hij : i ≠ j) : f (function.update v i (v j)) = 0 := f.map_eq_zero_of_eq _ (by rw [function.update_same, function.update_noteq hij.symm]) hij -lemma map_update_update {i j : ι} (hij : i ≠ j) (m : M) : +lemma map_update_update [decidable_eq ι] {i j : ι} (hij : i ≠ j) (m : M) : f (function.update (function.update v i m) j m) = 0 := f.map_eq_zero_of_eq _ (by rw [function.update_same, function.update_noteq hij, function.update_same]) hij -lemma map_swap_add {i j : ι} (hij : i ≠ j) : +lemma map_swap_add [decidable_eq ι] {i j : ι} (hij : i ≠ j) : f (v ∘ equiv.swap i j) + f v = 0 := begin rw equiv.comp_swap_eq_update, @@ -474,14 +475,14 @@ begin function.update_comm hij.symm (v i) (v i) v], end -lemma map_add_swap {i j : ι} (hij : i ≠ j) : +lemma map_add_swap [decidable_eq ι] {i j : ι} (hij : i ≠ j) : f v + f (v ∘ equiv.swap i j) = 0 := by { rw add_comm, exact f.map_swap_add v hij } -lemma map_swap {i j : ι} (hij : i ≠ j) : g (v ∘ equiv.swap i j) = - g v := +lemma map_swap [decidable_eq ι] {i j : ι} (hij : i ≠ j) : g (v ∘ equiv.swap i j) = - g v := eq_neg_of_add_eq_zero_left $ g.map_swap_add v hij -lemma map_perm [fintype ι] (v : ι → M) (σ : equiv.perm ι) : +lemma map_perm [decidable_eq ι] [fintype ι] (v : ι → M) (σ : equiv.perm ι) : g (v ∘ σ) = σ.sign • g v := begin apply equiv.perm.swap_induction_on' σ, @@ -490,7 +491,7 @@ begin simpa [g.map_swap (v ∘ s) hxy, equiv.perm.sign_swap hxy] using hI, } end -lemma map_congr_perm [fintype ι] (σ : equiv.perm ι) : +lemma map_congr_perm [decidable_eq ι] [fintype ι] (σ : equiv.perm ι) : g v = σ.sign • g (v ∘ σ) := by { rw [g.map_perm, smul_smul], simp } @@ -541,7 +542,7 @@ def dom_dom_congr_equiv (σ : ι ≃ ι') : f.dom_dom_congr σ = 0 ↔ f = 0 := (dom_dom_congr_equiv σ : alternating_map R M N ι ≃+ alternating_map R M N ι').map_eq_zero_iff -lemma dom_dom_congr_perm [fintype ι] (σ : equiv.perm ι) : +lemma dom_dom_congr_perm [fintype ι] [decidable_eq ι] (σ : equiv.perm ι) : g.dom_dom_congr σ = σ.sign • g := alternating_map.ext $ λ v, g.map_perm v σ @@ -561,6 +562,7 @@ lemma map_linear_dependent f v = 0 := begin obtain ⟨s, g, h, i, hi, hz⟩ := not_linear_independent_iff.mp h, + letI := classical.dec_eq ι, suffices : f (update v i (g i • v i)) = 0, { rw [f.map_smul, function.update_eq_self, smul_eq_zero] at this, exact or.resolve_left this hz, }, @@ -597,7 +599,7 @@ namespace multilinear_map open equiv -variables [fintype ι] +variables [fintype ι] [decidable_eq ι] private lemma alternization_map_eq_zero_of_eq_aux (m : multilinear_map R (λ i : ι, M) N') @@ -651,7 +653,7 @@ namespace alternating_map /-- Alternatizing a multilinear map that is already alternating results in a scale factor of `n!`, where `n` is the number of inputs. -/ -lemma coe_alternatization [fintype ι] (a : alternating_map R M N' ι) : +lemma coe_alternatization [decidable_eq ι] [fintype ι] (a : alternating_map R M N' ι) : (↑a : multilinear_map R (λ ι, M) N').alternatization = nat.factorial (fintype.card ι) • a := begin apply alternating_map.coe_injective, @@ -664,7 +666,7 @@ end alternating_map namespace linear_map -variables {N'₂ : Type*} [add_comm_group N'₂] [module R N'₂] [fintype ι] +variables {N'₂ : Type*} [add_comm_group N'₂] [module R N'₂] [decidable_eq ι] [fintype ι] /-- Composition with a linear map before and after alternatization are equivalent. -/ lemma comp_multilinear_map_alternatization (g : N' →ₗ[R] N'₂) @@ -679,7 +681,7 @@ section coprod open_locale big_operators open_locale tensor_product -variables {ιa ιb : Type*} [decidable_eq ιa] [decidable_eq ιb] [fintype ιa] [fintype ιb] +variables {ιa ιb : Type*}[fintype ιa] [fintype ιb] variables {R' : Type*} {Mᵢ N₁ N₂ : Type*} @@ -705,6 +707,7 @@ end equiv.perm namespace alternating_map open equiv +variables [decidable_eq ιa] [decidable_eq ιb] /-- summand used in `alternating_map.dom_coprod` -/ def dom_coprod.summand @@ -869,7 +872,7 @@ end alternating_map open equiv /-- A helper lemma for `multilinear_map.dom_coprod_alternization`. -/ -lemma multilinear_map.dom_coprod_alternization_coe +lemma multilinear_map.dom_coprod_alternization_coe [decidable_eq ιa] [decidable_eq ιb] (a : multilinear_map R' (λ _ : ιa, Mᵢ) N₁) (b : multilinear_map R' (λ _ : ιb, Mᵢ) N₂) : multilinear_map.dom_coprod ↑a.alternatization ↑b.alternatization = ∑ (σa : perm ιa) (σb : perm ιb), σa.sign • σb.sign • @@ -885,7 +888,7 @@ open alternating_map /-- Computing the `multilinear_map.alternatization` of the `multilinear_map.dom_coprod` is the same as computing the `alternating_map.dom_coprod` of the `multilinear_map.alternatization`s. -/ -lemma multilinear_map.dom_coprod_alternization +lemma multilinear_map.dom_coprod_alternization [decidable_eq ιa] [decidable_eq ιb] (a : multilinear_map R' (λ _ : ιa, Mᵢ) N₁) (b : multilinear_map R' (λ _ : ιb, Mᵢ) N₂) : (multilinear_map.dom_coprod a b).alternatization = a.alternatization.dom_coprod b.alternatization := @@ -932,7 +935,7 @@ end /-- Taking the `multilinear_map.alternatization` of the `multilinear_map.dom_coprod` of two `alternating_map`s gives a scaled version of the `alternating_map.coprod` of those maps. -/ -lemma multilinear_map.dom_coprod_alternization_eq +lemma multilinear_map.dom_coprod_alternization_eq [decidable_eq ιa] [decidable_eq ιb] (a : alternating_map R' Mᵢ N₁ ιa) (b : alternating_map R' Mᵢ N₂ ιb) : (multilinear_map.dom_coprod a b : multilinear_map R' (λ _ : ιa ⊕ ιb, Mᵢ) (N₁ ⊗ N₂)) .alternatization = @@ -960,6 +963,7 @@ are distinct basis vectors. -/ lemma basis.ext_alternating {f g : alternating_map R' N₁ N₂ ι} (e : basis ι₁ R' N₁) (h : ∀ v : ι → ι₁, function.injective v → f (λ i, e (v i)) = g (λ i, e (v i))) : f = g := begin + classical, refine alternating_map.coe_multilinear_map_injective (basis.ext_multilinear e $ λ v, _), by_cases hi : function.injective v, { exact h v hi }, diff --git a/src/linear_algebra/determinant.lean b/src/linear_algebra/determinant.lean index 1623a1ec203b5..8f4a7c9dbf6f8 100644 --- a/src/linear_algebra/determinant.lean +++ b/src/linear_algebra/determinant.lean @@ -452,12 +452,14 @@ multilinear map. -/ def basis.det : alternating_map R M R ι := { to_fun := λ v, det (e.to_matrix v), map_add' := begin - intros v i x y, - simp only [e.to_matrix_update, linear_equiv.map_add], - apply det_update_column_add + intros inst v i x y, + cases subsingleton.elim inst ‹_›, + simp only [e.to_matrix_update, linear_equiv.map_add, finsupp.coe_add], + exact det_update_column_add _ _ _ _, end, map_smul' := begin - intros u i c x, + intros inst u i c x, + cases subsingleton.elim inst ‹_›, simp only [e.to_matrix_update, algebra.id.smul_eq_mul, linear_equiv.map_smul], apply det_update_column_smul end, @@ -516,13 +518,16 @@ begin simp [alternating_map.map_perm, basis.det_self] end -@[simp] lemma alternating_map.map_basis_eq_zero_iff {ι : Type*} [decidable_eq ι] [finite ι] +@[simp] lemma alternating_map.map_basis_eq_zero_iff {ι : Type*} [finite ι] (e : basis ι R M) (f : alternating_map R M R ι) : f e = 0 ↔ f = 0 := -⟨λ h, by { casesI nonempty_fintype ι, simpa [h] using f.eq_smul_basis_det e }, - λ h, h.symm ▸ alternating_map.zero_apply _⟩ +⟨λ h, begin + casesI nonempty_fintype ι, + letI := classical.dec_eq ι, + simpa [h] using f.eq_smul_basis_det e +end, λ h, h.symm ▸ alternating_map.zero_apply _⟩ -lemma alternating_map.map_basis_ne_zero_iff {ι : Type*} [decidable_eq ι] [finite ι] +lemma alternating_map.map_basis_ne_zero_iff {ι : Type*} [finite ι] (e : basis ι R M) (f : alternating_map R M R ι) : f e ≠ 0 ↔ f ≠ 0 := not_congr $ f.map_basis_eq_zero_iff e diff --git a/src/linear_algebra/exterior_algebra/of_alternating.lean b/src/linear_algebra/exterior_algebra/of_alternating.lean index 2ad764e901130..0aaeb47af7172 100644 --- a/src/linear_algebra/exterior_algebra/of_alternating.lean +++ b/src/linear_algebra/exterior_algebra/of_alternating.lean @@ -28,7 +28,7 @@ variables [comm_ring R] [add_comm_group M] [add_comm_group N] [add_comm_group N' variables [module R M] [module R N] [module R N'] -- This instance can't be found where it's needed if we don't remind lean that it exists. -instance alternating_map.module_add_comm_group {ι : Type*} [decidable_eq ι] : +instance alternating_map.module_add_comm_group {ι : Type*} : module R (alternating_map R M N ι) := by apply_instance diff --git a/src/linear_algebra/multilinear/basic.lean b/src/linear_algebra/multilinear/basic.lean index f2df2f37429c2..ce8c160726d99 100644 --- a/src/linear_algebra/multilinear/basic.lean +++ b/src/linear_algebra/multilinear/basic.lean @@ -45,11 +45,29 @@ in linear functions), called respectively `multilinear_curry_left_equiv` and Expressing that a map is linear along the `i`-th coordinate when all other coordinates are fixed can be done in two (equivalent) different ways: + * fixing a vector `m : Π(j : ι - i), M₁ j.val`, and then choosing separately the `i`-th coordinate * fixing a vector `m : Πj, M₁ j`, and then modifying its `i`-th coordinate + The second way is more artificial as the value of `m` at `i` is not relevant, but it has the advantage of avoiding subtype inclusion issues. This is the definition we use, based on `function.update` that allows to change the value of `m` at `i`. + +Note that the use of `function.update` requires a `decidable_eq ι` term to appear somewhere in the +statement of `multilinear_map.map_add'` and `multilinear_map.map_smul'`. Three possible choices +are: + +1. Requiring `decidable_eq ι` as an argument to `multilinear_map` (as we did originally). +2. Using `classical.dec_eq ι` in the statement of `map_add'` and `map_smul'`. +3. Quantifying over all possible `decidable_eq ι` instances in the statement of `map_add'` and + `map_smul'`. + +Option 1 works fine, but puts unecessary constraints on the user (the zero map certainly does not +need decidability). Option 2 looks great at first, but in the common case when `ι = fin n` it +introduces non-defeq decidability instance diamonds within the context of proving `map_add'` and +`map_smul'`, of the form `fin.decidable_eq n = classical.dec_eq (fin n)`. Option 3 of course does +something similar, but of the form `fin.decidable_eq n = _inst`, which is much easier to clean up +since `_inst` is a free variable and so the equality can just be substituted. -/ open function fin set @@ -58,17 +76,16 @@ open_locale big_operators universes u v v' v₁ v₂ v₃ w u' variables {R : Type u} {ι : Type u'} {n : ℕ} {M : fin n.succ → Type v} {M₁ : ι → Type v₁} {M₂ : Type v₂} {M₃ : Type v₃} {M' : Type v'} -[decidable_eq ι] /-- Multilinear maps over the ring `R`, from `Πi, M₁ i` to `M₂` where `M₁ i` and `M₂` are modules over `R`. -/ structure multilinear_map (R : Type u) {ι : Type u'} (M₁ : ι → Type v) (M₂ : Type w) - [decidable_eq ι] [semiring R] [∀i, add_comm_monoid (M₁ i)] [add_comm_monoid M₂] + [semiring R] [∀i, add_comm_monoid (M₁ i)] [add_comm_monoid M₂] [∀i, module R (M₁ i)] [module R M₂] := (to_fun : (Πi, M₁ i) → M₂) -(map_add' : ∀(m : Πi, M₁ i) (i : ι) (x y : M₁ i), +(map_add' : ∀ [decidable_eq ι] (m : Πi, M₁ i) (i : ι) (x y : M₁ i), by exactI to_fun (update m i (x + y)) = to_fun (update m i x) + to_fun (update m i y)) -(map_smul' : ∀(m : Πi, M₁ i) (i : ι) (c : R) (x : M₁ i), +(map_smul' : ∀ [decidable_eq ι] (m : Πi, M₁ i) (i : ι) (c : R) (x : M₁ i), by exactI to_fun (update m i (c • x)) = c • to_fun (update m i x)) namespace multilinear_map @@ -114,21 +131,22 @@ theorem ext_iff {f g : multilinear_map R M₁ M₂} : f = g ↔ ∀ x, f x = g x (⟨f, h₁, h₂⟩ : multilinear_map R M₁ M₂) = f := by { ext, refl, } -@[simp] protected lemma map_add (m : Πi, M₁ i) (i : ι) (x y : M₁ i) : +@[simp] protected lemma map_add [decidable_eq ι] (m : Πi, M₁ i) (i : ι) (x y : M₁ i) : f (update m i (x + y)) = f (update m i x) + f (update m i y) := f.map_add' m i x y -@[simp] protected lemma map_smul (m : Πi, M₁ i) (i : ι) (c : R) (x : M₁ i) : +@[simp] protected lemma map_smul [decidable_eq ι] (m : Πi, M₁ i) (i : ι) (c : R) (x : M₁ i) : f (update m i (c • x)) = c • f (update m i x) := f.map_smul' m i c x lemma map_coord_zero {m : Πi, M₁ i} (i : ι) (h : m i = 0) : f m = 0 := begin + classical, have : (0 : R) • (0 : M₁ i) = 0, by simp, rw [← update_eq_self i m, h, ← this, f.map_smul, zero_smul] end -@[simp] lemma map_update_zero (m : Πi, M₁ i) (i : ι) : f (update m i 0) = 0 := +@[simp] lemma map_update_zero [decidable_eq ι] (m : Πi, M₁ i) (i : ι) : f (update m i 0) = 0 := f.map_coord_zero i (update_same i 0 m) @[simp] lemma map_zero [nonempty ι] : f 0 = 0 := @@ -139,12 +157,12 @@ end instance : has_add (multilinear_map R M₁ M₂) := ⟨λf f', ⟨λx, f x + f' x, λm i x y, by simp [add_left_comm, add_assoc], - λm i c x, by simp [smul_add]⟩⟩ + λ _ m i c x, by simp [smul_add]⟩⟩ @[simp] lemma add_apply (m : Πi, M₁ i) : (f + f') m = f m + f' m := rfl instance : has_zero (multilinear_map R M₁ M₂) := -⟨⟨λ _, 0, λm i x y, by simp, λm i c x, by simp⟩⟩ +⟨⟨λ _, 0, λ _ m i x y, by simp, λ _ m i c x, by simp⟩⟩ instance : inhabited (multilinear_map R M₁ M₂) := ⟨0⟩ @@ -155,7 +173,7 @@ variables {R' A : Type*} [monoid R'] [semiring A] [Π i, module A (M₁ i)] [distrib_mul_action R' M₂] [module A M₂] [smul_comm_class A R' M₂] instance : has_smul R' (multilinear_map A M₁ M₂) := ⟨λ c f, - ⟨λ m, c • f m, λm i x y, by simp [smul_add], λl i x d, by simp [←smul_comm x c] ⟩⟩ + ⟨λ m, c • f m, λ _ m i x y, by simp [smul_add], λ _ l i x d, by simp [←smul_comm x c] ⟩⟩ @[simp] lemma smul_apply (f : multilinear_map A M₁ M₂) (c : R') (m : Πi, M₁ i) : (c • f) m = c • f m := rfl @@ -179,7 +197,7 @@ end /-- If `f` is a multilinear map, then `f.to_linear_map m i` is the linear map obtained by fixing all coordinates but `i` equal to those of `m`, and varying the `i`-th coordinate. -/ -@[simps] def to_linear_map (m : Πi, M₁ i) (i : ι) : M₁ i →ₗ[R] M₂ := +@[simps] def to_linear_map [decidable_eq ι] (m : Πi, M₁ i) (i : ι) : M₁ i →ₗ[R] M₂ := { to_fun := λx, f (update m i x), map_add' := λx y, by simp, map_smul' := λc x, by simp } @@ -188,8 +206,8 @@ coordinates but `i` equal to those of `m`, and varying the `i`-th coordinate. -/ def prod (f : multilinear_map R M₁ M₂) (g : multilinear_map R M₁ M₃) : multilinear_map R M₁ (M₂ × M₃) := { to_fun := λ m, (f m, g m), - map_add' := λ m i x y, by simp, - map_smul' := λ m i c x, by simp } + map_add' := λ _ m i x y, by simp, + map_smul' := λ _ m i c x, by simp } /-- Combine a family of multilinear maps with the same domain and codomains `M' i` into a multilinear map taking values in the space of functions `Π i, M' i`. -/ @@ -197,8 +215,8 @@ multilinear map taking values in the space of functions `Π i, M' i`. -/ [Π i, module R (M' i)] (f : Π i, multilinear_map R M₁ (M' i)) : multilinear_map R M₁ (Π i, M' i) := { to_fun := λ m i, f i m, - map_add' := λ m i x y, funext $ λ j, (f j).map_add _ _ _ _, - map_smul' := λ m i c x, funext $ λ j, (f j).map_smul _ _ _ _ } + map_add' := λ _ m i x y, by exactI (funext $ λ j, (f j).map_add _ _ _ _), + map_smul' := λ _ m i c x, by exactI (funext $ λ j, (f j).map_smul _ _ _ _) } section variables (R M₂) @@ -208,9 +226,9 @@ variables (R M₂) @[simps] def of_subsingleton [subsingleton ι] (i' : ι) : multilinear_map R (λ _ : ι, M₂) M₂ := { to_fun := function.eval i', - map_add' := λ m i x y, by + map_add' := λ _ m i x y, by { rw subsingleton.elim i i', simp only [function.eval, function.update_same], }, - map_smul' := λ m i r x, by + map_smul' := λ _ m i r x, by { rw subsingleton.elim i i', simp only [function.eval, function.update_same], } } variables (M₁) {M₂} @@ -219,8 +237,8 @@ variables (M₁) {M₂} @[simps {fully_applied := ff}] def const_of_is_empty [is_empty ι] (m : M₂) : multilinear_map R M₁ M₂ := { to_fun := function.const _ m, - map_add' := λ m, is_empty_elim, - map_smul' := λ m, is_empty_elim } + map_add' := λ _ m, is_empty_elim, + map_smul' := λ _ m, is_empty_elim } end @@ -233,9 +251,9 @@ def restr {k n : ℕ} (f : multilinear_map R (λ i : fin n, M') M₂) (s : finse (hk : s.card = k) (z : M') : multilinear_map R (λ i : fin k, M') M₂ := { to_fun := λ v, f (λ j, if h : j ∈ s then v ((s.order_iso_of_fin hk).symm ⟨j, h⟩) else z), - map_add' := λ v i x y, + map_add' := λ _ v i x y, by { erw [dite_comp_equiv_update, dite_comp_equiv_update, dite_comp_equiv_update], simp }, - map_smul' := λ v i c x, by { erw [dite_comp_equiv_update, dite_comp_equiv_update], simp } } + map_smul' := λ _ v i c x, by { erw [dite_comp_equiv_update, dite_comp_equiv_update], simp } } variable {R} /-- In the specific case of multilinear maps on spaces indexed by `fin (n+1)`, where one can build @@ -278,14 +296,16 @@ then `g (f₁ m₁, ..., fₙ mₙ)` is again a multilinear map, that we call def comp_linear_map (g : multilinear_map R M₁' M₂) (f : Π i, M₁ i →ₗ[R] M₁' i) : multilinear_map R M₁ M₂ := { to_fun := λ m, g $ λ i, f i (m i), - map_add' := λ m i x y, - have ∀ j z, f j (update m i z j) = update (λ k, f k (m k)) i (f i z) j := + map_add' := λ _ m i x y, by + { resetI, + have : ∀ j z, f j (update m i z j) = update (λ k, f k (m k)) i (f i z) j := λ j z, function.apply_update (λ k, f k) _ _ _ _, - by simp [this], - map_smul' := λ m i c x, - have ∀ j z, f j (update m i z j) = update (λ k, f k (m k)) i (f i z) j := + by simp [this] }, + map_smul' := λ _ m i c x, by + { resetI, + have : ∀ j z, f j (update m i z j) = update (λ k, f k (m k)) i (f i z) j := λ j z, function.apply_update (λ k, f k) _ _ _ _, - by simp [this] } + by simp [this] } } @[simp] lemma comp_linear_map_apply (g : multilinear_map R M₁' M₂) (f : Π i, M₁ i →ₗ[R] M₁' i) (m : Π i, M₁ i) : @@ -335,7 +355,7 @@ the image under a multilinear map `f` is the sum of `f (s.piecewise m m')` along `t`. This is mainly an auxiliary statement to prove the result when `t = univ`, given in `map_add_univ`, although it can be useful in its own right as it does not require the index set `ι` to be finite.-/ -lemma map_piecewise_add (m m' : Πi, M₁ i) (t : finset ι) : +lemma map_piecewise_add [decidable_eq ι] (m m' : Πi, M₁ i) (t : finset ι) : f (t.piecewise (m + m') m') = ∑ s in t.powerset, f (s.piecewise m m') := begin revert m', @@ -367,7 +387,7 @@ end /-- Additivity of a multilinear map along all coordinates at the same time, writing `f (m + m')` as the sum of `f (s.piecewise m m')` over all sets `s`. -/ -lemma map_add_univ [fintype ι] (m m' : Πi, M₁ i) : +lemma map_add_univ [decidable_eq ι] [fintype ι] (m m' : Πi, M₁ i) : f (m + m') = ∑ s : finset ι, f (s.piecewise m m') := by simpa using f.map_piecewise_add m m' finset.univ @@ -375,7 +395,6 @@ section apply_sum variables {α : ι → Type*} (g : Π i, α i → M₁ i) (A : Π i, finset (α i)) -open_locale classical open fintype finset /-- If `f` is multilinear, then `f (Σ_{j₁ ∈ A₁} g₁ j₁, ..., Σ_{jₙ ∈ Aₙ} gₙ jₙ)` is the sum of @@ -383,9 +402,10 @@ open fintype finset `r n ∈ Aₙ`. This follows from multilinearity by expanding successively with respect to each coordinate. Here, we give an auxiliary statement tailored for an inductive proof. Use instead `map_sum_finset`. -/ -lemma map_sum_finset_aux [fintype ι] {n : ℕ} (h : ∑ i, (A i).card = n) : +lemma map_sum_finset_aux [decidable_eq ι] [fintype ι] {n : ℕ} (h : ∑ i, (A i).card = n) : f (λ i, ∑ j in A i, g i j) = ∑ r in pi_finset A, f (λ i, g i (r i)) := begin + letI := λ i, classical.dec_eq (α i), induction n using nat.strong_induction_on with n IH generalizing A, -- If one of the sets is empty, then all the sums are zero by_cases Ai_empty : ∃ i, A i = ∅, @@ -519,20 +539,22 @@ end `f (g₁ (r 1), ..., gₙ (r n))` where `r` ranges over all functions with `r 1 ∈ A₁`, ..., `r n ∈ Aₙ`. This follows from multilinearity by expanding successively with respect to each coordinate. -/ -lemma map_sum_finset [fintype ι] : +lemma map_sum_finset [decidable_eq ι] [fintype ι] : f (λ i, ∑ j in A i, g i j) = ∑ r in pi_finset A, f (λ i, g i (r i)) := f.map_sum_finset_aux _ _ rfl /-- If `f` is multilinear, then `f (Σ_{j₁} g₁ j₁, ..., Σ_{jₙ} gₙ jₙ)` is the sum of `f (g₁ (r 1), ..., gₙ (r n))` where `r` ranges over all functions `r`. This follows from multilinearity by expanding successively with respect to each coordinate. -/ -lemma map_sum [fintype ι] [∀ i, fintype (α i)] : +lemma map_sum [decidable_eq ι] [fintype ι] [∀ i, fintype (α i)] : f (λ i, ∑ j, g i j) = ∑ r : Π i, α i, f (λ i, g i (r i)) := f.map_sum_finset g (λ i, finset.univ) -lemma map_update_sum {α : Type*} (t : finset α) (i : ι) (g : α → M₁ i) (m : Π i, M₁ i): +lemma map_update_sum {α : Type*} [decidable_eq ι] (t : finset α) (i : ι) (g : α → M₁ i) + (m : Π i, M₁ i) : f (update m i (∑ a in t, g a)) = ∑ a in t, f (update m i (g a)) := begin + classical, induction t using finset.induction with a t has ih h, { simp }, { simp [finset.sum_insert has, ih] } @@ -547,8 +569,8 @@ This is the multilinear version of `linear_map.cod_restrict`. -/ def cod_restrict (f : multilinear_map R M₁ M₂) (p : submodule R M₂) (h : ∀ v, f v ∈ p) : multilinear_map R M₁ p := { to_fun := λ v, ⟨f v, h v⟩, - map_add' := λ v i x y, subtype.ext $ multilinear_map.map_add _ _ _ _ _, - map_smul' := λ v i c x, subtype.ext $ multilinear_map.map_smul _ _ _ _ _ } + map_add' := λ _ v i x y, subtype.ext $ by exactI multilinear_map.map_add _ _ _ _ _, + map_smul' := λ _ v i c x, subtype.ext $ by exactI multilinear_map.map_smul _ _ _ _ _ } section restrict_scalar @@ -559,8 +581,8 @@ variables (R) {A : Type*} [semiring A] [has_smul R A] [Π (i : ι), module A (M and their actions on all involved modules agree with the action of `R` on `A`. -/ def restrict_scalars (f : multilinear_map A M₁ M₂) : multilinear_map R M₁ M₂ := { to_fun := f, - map_add' := f.map_add, - map_smul' := λ m i, (f.to_linear_map m i).map_smul_of_tower } + map_add' := λ _, by exactI f.map_add, + map_smul' := λ _ m i, by exactI (f.to_linear_map m i).map_smul_of_tower } @[simp] lemma coe_restrict_scalars (f : multilinear_map A M₁ M₂) : ⇑(f.restrict_scalars R) = f := rfl @@ -570,7 +592,7 @@ end restrict_scalar section -variables {ι₁ ι₂ ι₃ : Type*} [decidable_eq ι₁] [decidable_eq ι₂] [decidable_eq ι₃] +variables {ι₁ ι₂ ι₃ : Type*} /-- Transfer the arguments to a map along an equivalence between argument indices. @@ -580,8 +602,12 @@ domain of the domain. -/ def dom_dom_congr (σ : ι₁ ≃ ι₂) (m : multilinear_map R (λ i : ι₁, M₂) M₃) : multilinear_map R (λ i : ι₂, M₂) M₃ := { to_fun := λ v, m (λ i, v (σ i)), - map_add' := λ v i a b, by { simp_rw function.update_apply_equiv_apply v, rw m.map_add, }, - map_smul' := λ v i a b, by { simp_rw function.update_apply_equiv_apply v, rw m.map_smul, }, } + map_add' := λ _ v i a b, by + { resetI, letI := σ.injective.decidable_eq, + simp_rw function.update_apply_equiv_apply v, rw m.map_add, }, + map_smul' := λ _ v i a b, by + { resetI, letI := σ.injective.decidable_eq, + simp_rw function.update_apply_equiv_apply v, rw m.map_smul, }, } lemma dom_dom_congr_trans (σ₁ : ι₁ ≃ ι₂) (σ₂ : ι₂ ≃ ι₃) (m : multilinear_map R (λ i : ι₁, M₂) M₃) : m.dom_dom_congr (σ₁.trans σ₂) = (m.dom_dom_congr σ₁).dom_dom_congr σ₂ := rfl @@ -647,7 +673,7 @@ lemma comp_multilinear_map_cod_restrict (g : M₂ →ₗ[R] M₃) (f : multiline (g.comp_multilinear_map f).cod_restrict p (λ v, h (f v)):= multilinear_map.ext $ λ v, rfl -variables {ι₁ ι₂ : Type*} [decidable_eq ι₁] [decidable_eq ι₂] +variables {ι₁ ι₂ : Type*} @[simp] lemma comp_multilinear_map_dom_dom_congr (σ : ι₁ ≃ ι₂) (g : M₂ →ₗ[R] M₃) (f : multilinear_map R (λ i : ι₁, M') M₂) : @@ -668,7 +694,7 @@ variables [comm_semiring R] [∀i, add_comm_monoid (M₁ i)] [∀i, add_comm_mon map is multiplied by `∏ i in s, c i`. This is mainly an auxiliary statement to prove the result when `s = univ`, given in `map_smul_univ`, although it can be useful in its own right as it does not require the index set `ι` to be finite. -/ -lemma map_piecewise_smul (c : ι → R) (m : Πi, M₁ i) (s : finset ι) : +lemma map_piecewise_smul [decidable_eq ι] (c : ι → R) (m : Πi, M₁ i) (s : finset ι) : f (s.piecewise (λi, c i • m i) m) = (∏ i in s, c i) • f m := begin refine s.induction_on (by simp) _, @@ -687,9 +713,10 @@ end writing `f (λi, c i • m i)` as `(∏ i, c i) • f m`. -/ lemma map_smul_univ [fintype ι] (c : ι → R) (m : Πi, M₁ i) : f (λi, c i • m i) = (∏ i, c i) • f m := -by simpa using map_piecewise_smul f c m finset.univ +by {classical, simpa using map_piecewise_smul f c m finset.univ} -@[simp] lemma map_update_smul [fintype ι] (m : Πi, M₁ i) (i : ι) (c : R) (x : M₁ i) : +@[simp] lemma map_update_smul [decidable_eq ι] [fintype ι] (m : Πi, M₁ i) (i : ι) (c : R) + (x : M₁ i) : f (update (c • m) i x) = c^(fintype.card ι - 1) • f (update m i x) := begin have : f ((finset.univ.erase i).piecewise (c • update m i x) (update m i x)) @@ -729,42 +756,49 @@ variables (M₂ M₃ R' A) /-- `multilinear_map.dom_dom_congr` as a `linear_equiv`. -/ @[simps apply symm_apply] -def dom_dom_congr_linear_equiv {ι₁ ι₂} [decidable_eq ι₁] [decidable_eq ι₂] (σ : ι₁ ≃ ι₂) : +def dom_dom_congr_linear_equiv {ι₁ ι₂} (σ : ι₁ ≃ ι₂) : multilinear_map A (λ i : ι₁, M₂) M₃ ≃ₗ[R'] multilinear_map A (λ i : ι₂, M₂) M₃ := { map_smul' := λ c f, by { ext, simp }, .. (dom_dom_congr_equiv σ : multilinear_map A (λ i : ι₁, M₂) M₃ ≃+ multilinear_map A (λ i : ι₂, M₂) M₃) } variables (R M₁) - /-- The dependent version of `multilinear_map.dom_dom_congr_linear_equiv`. -/ @[simps apply symm_apply] -def dom_dom_congr_linear_equiv' {ι' : Type*} [decidable_eq ι'] (σ : ι ≃ ι') : +def dom_dom_congr_linear_equiv' {ι' : Type*} (σ : ι ≃ ι') : multilinear_map R M₁ M₂ ≃ₗ[R] multilinear_map R (λ i, M₁ (σ.symm i)) M₂ := { to_fun := λ f, { to_fun := f ∘ (σ.Pi_congr_left' M₁).symm, - map_add' := λ m i, + map_add' := λ _ m i, begin + resetI, + letI := σ.decidable_eq, rw ← σ.apply_symm_apply i, intros x y, simp only [comp_app, Pi_congr_left'_symm_update, f.map_add], end, - map_smul' := λ m i c, + map_smul' := λ _ m i c, begin + resetI, + letI := σ.decidable_eq, rw ← σ.apply_symm_apply i, intros x, simp only [comp_app, Pi_congr_left'_symm_update, f.map_smul], end, }, inv_fun := λ f, { to_fun := f ∘ (σ.Pi_congr_left' M₁), - map_add' := λ m i, + map_add' := λ _ m i, begin + resetI, + letI := σ.symm.decidable_eq, rw ← σ.symm_apply_apply i, intros x y, simp only [comp_app, Pi_congr_left'_update, f.map_add], end, - map_smul' := λ m i c, + map_smul' := λ _ m i c, begin + resetI, + letI := σ.symm.decidable_eq, rw ← σ.symm_apply_apply i, intros x, simp only [comp_app, Pi_congr_left'_update, f.map_smul], @@ -821,7 +855,8 @@ protected def mk_pi_algebra_fin : multilinear_map R (λ i : fin n, A) A := { to_fun := λ m, (list.of_fn m).prod, map_add' := begin - intros m i x y, + intros dec m i x y, + rw subsingleton.elim dec (by apply_instance), have : (list.fin_range n).index_of i < n, by simpa using list.index_of_lt_length.2 (list.mem_fin_range i), simp [list.of_fn_eq_map, (list.nodup_fin_range n).map_update, list.prod_update_nth, add_mul, @@ -829,7 +864,8 @@ protected def mk_pi_algebra_fin : multilinear_map R (λ i : fin n, A) A := end, map_smul' := begin - intros m i c x, + intros dec m i c x, + rw subsingleton.elim dec (by apply_instance), have : (list.fin_range n).index_of i < n, by simpa using list.index_of_lt_length.2 (list.mem_fin_range i), simp [list.of_fn_eq_map, (list.nodup_fin_range n).map_update, list.prod_update_nth, this] @@ -904,32 +940,33 @@ variables [semiring R] [∀i, add_comm_monoid (M₁ i)] [add_comm_group M₂] (f g : multilinear_map R M₁ M₂) instance : has_neg (multilinear_map R M₁ M₂) := -⟨λ f, ⟨λ m, - f m, λm i x y, by simp [add_comm], λm i c x, by simp⟩⟩ +⟨λ f, ⟨λ m, - f m, λ _ m i x y, by simp [add_comm], λ _ m i c x, by simp⟩⟩ @[simp] lemma neg_apply (m : Πi, M₁ i) : (-f) m = - (f m) := rfl instance : has_sub (multilinear_map R M₁ M₂) := ⟨λ f g, ⟨λ m, f m - g m, - λ m i x y, by { simp only [multilinear_map.map_add, sub_eq_add_neg, neg_add], cc }, - λ m i c x, by { simp only [multilinear_map.map_smul, smul_sub] }⟩⟩ + λ _ m i x y, by { simp only [multilinear_map.map_add, sub_eq_add_neg, neg_add], cc }, + λ _ m i c x, by { simp only [multilinear_map.map_smul, smul_sub] }⟩⟩ @[simp] lemma sub_apply (m : Πi, M₁ i) : (f - g) m = f m - g m := rfl instance : add_comm_group (multilinear_map R M₁ M₂) := -by refine { zero := (0 : multilinear_map R M₁ M₂), add := (+), neg := has_neg.neg, sub := has_sub.sub, - sub_eq_add_neg := _, - nsmul := λ n f, ⟨λ m, n • f m, λm i x y, by simp [smul_add], λl i x d, by simp [←smul_comm x n] ⟩, - zsmul := λ n f, ⟨λ m, n • f m, λm i x y, by simp [smul_add], λl i x d, by simp [←smul_comm x n] ⟩, - zsmul_zero' := _, - zsmul_succ' := _, - zsmul_neg' := _, - .. multilinear_map.add_comm_monoid, .. }; -intros; ext; simp [add_comm, add_left_comm, sub_eq_add_neg, add_smul, nat.succ_eq_add_one] + add_left_neg := λ a, multilinear_map.ext $ λ v, add_left_neg _, + sub_eq_add_neg := λ a b, multilinear_map.ext $ λ v, sub_eq_add_neg _ _, + zsmul := λ n f, + { to_fun := λ m, n • f m, + map_add' := λ _ m i x y, by simp [smul_add], + map_smul' := λ _ l i x d, by simp [←smul_comm x n]}, + zsmul_zero' := λ a, multilinear_map.ext $ λ v, add_comm_group.zsmul_zero' _, + zsmul_succ' := λ z a, multilinear_map.ext $ λ v, add_comm_group.zsmul_succ' _ _, + zsmul_neg' := λ z a, multilinear_map.ext $ λ v, add_comm_group.zsmul_neg' _ _, + .. multilinear_map.add_comm_monoid } end range_add_comm_group @@ -939,12 +976,12 @@ variables [semiring R] [∀i, add_comm_group (M₁ i)] [add_comm_group M₂] [∀i, module R (M₁ i)] [module R M₂] (f : multilinear_map R M₁ M₂) -@[simp] lemma map_neg (m : Πi, M₁ i) (i : ι) (x : M₁ i) : +@[simp] lemma map_neg [decidable_eq ι] (m : Πi, M₁ i) (i : ι) (x : M₁ i) : f (update m i (-x)) = -f (update m i x) := eq_neg_of_add_eq_zero_left $ by rw [←multilinear_map.map_add, add_left_neg, f.map_coord_zero i (update_same i 0 m)] -@[simp] lemma map_sub (m : Πi, M₁ i) (i : ι) (x y : M₁ i) : +@[simp] lemma map_sub [decidable_eq ι] (m : Πi, M₁ i) (i : ι) (x y : M₁ i) : f (update m i (x - y)) = f (update m i x) - f (update m i y) := by rw [sub_eq_add_neg, sub_eq_add_neg, multilinear_map.map_add, map_neg] @@ -999,7 +1036,8 @@ def linear_map.uncurry_left (f : M 0 →ₗ[R] (multilinear_map R (λ(i : fin n), M i.succ) M₂)) : multilinear_map R M M₂ := { to_fun := λm, f (m 0) (tail m), - map_add' := λm i x y, begin + map_add' := λ dec m i x y, begin + rw subsingleton.elim dec (by apply_instance), by_cases h : i = 0, { subst i, rw [update_same, update_same, update_same, f.map_add, add_apply, @@ -1010,7 +1048,8 @@ def linear_map.uncurry_left assume x y, rw [tail_update_succ, multilinear_map.map_add, tail_update_succ, tail_update_succ] } end, - map_smul' := λm i c x, begin + map_smul' := λ dec m i c x, begin + rw subsingleton.elim dec (by apply_instance), by_cases h : i = 0, { subst i, rw [update_same, update_same, tail_update_zero, tail_update_zero, @@ -1032,9 +1071,9 @@ def multilinear_map.curry_left (f : multilinear_map R M M₂) : M 0 →ₗ[R] (multilinear_map R (λ(i : fin n), M i.succ) M₂) := { to_fun := λx, - { to_fun := λm, f (cons x m), - map_add' := λm i y y', by simp, - map_smul' := λm i y c, by simp }, + { to_fun := λ m, f (cons x m), + map_add' := λ dec m i y y', by { rw subsingleton.elim dec (by apply_instance), simp }, + map_smul' := λ dec m i y c, by { rw subsingleton.elim dec (by apply_instance), simp }, }, map_add' := λx y, by { ext m, exact cons_add f m x y }, map_smul' := λc x, by { ext m, exact cons_smul f m c x } } @@ -1085,7 +1124,8 @@ def multilinear_map.uncurry_right (f : (multilinear_map R (λ(i : fin n), M i.cast_succ) (M (last n) →ₗ[R] M₂))) : multilinear_map R M M₂ := { to_fun := λm, f (init m) (m (last n)), - map_add' := λm i x y, begin + map_add' := λ dec m i x y, begin + rw subsingleton.elim dec (by apply_instance), by_cases h : i.val < n, { have : last n ≠ i := ne.symm (ne_of_lt h), rw [update_noteq this, update_noteq this, update_noteq this], @@ -1100,7 +1140,8 @@ def multilinear_map.uncurry_right rw [init_update_last, init_update_last, init_update_last, update_same, update_same, update_same, linear_map.map_add] } end, - map_smul' := λm i c x, begin + map_smul' := λ dec m i c x, begin + rw subsingleton.elim dec (by apply_instance), by_cases h : i.val < n, { have : last n ≠ i := ne.symm (ne_of_lt h), rw [update_noteq this, update_noteq this], @@ -1128,13 +1169,15 @@ def multilinear_map.curry_right (f : multilinear_map R M M₂) : { to_fun := λx, f (snoc m x), map_add' := λx y, by rw f.snoc_add, map_smul' := λc x, by simp only [f.snoc_smul, ring_hom.id_apply] }, - map_add' := λm i x y, begin + map_add' := λ dec m i x y, begin + rw subsingleton.elim dec (by apply_instance), ext z, change f (snoc (update m i (x + y)) z) = f (snoc (update m i x) z) + f (snoc (update m i y) z), rw [snoc_update, snoc_update, snoc_update, f.map_add] end, - map_smul' := λm i c x, begin + map_smul' := λ dec m i c x, begin + rw subsingleton.elim dec (by apply_instance), ext z, change f (snoc (update m i (c • x)) z) = c • f (snoc (update m i x) z), rw [snoc_update, snoc_update, f.map_smul] @@ -1178,7 +1221,7 @@ def multilinear_curry_right_equiv : namespace multilinear_map -variables {ι' : Type*} [decidable_eq ι'] [decidable_eq (ι ⊕ ι')] {R M₂} +variables {ι' : Type*} {R M₂} /-- A multilinear map on `Π i : ι ⊕ ι', M'` defines a multilinear map on `Π i : ι, M'` taking values in the space of multilinear maps on `Π i : ι', M'`. -/ @@ -1186,12 +1229,18 @@ def curry_sum (f : multilinear_map R (λ x : ι ⊕ ι', M') M₂) : multilinear_map R (λ x : ι, M') (multilinear_map R (λ x : ι', M') M₂) := { to_fun := λ u, { to_fun := λ v, f (sum.elim u v), - map_add' := λ v i x y, by simp only [← sum.update_elim_inr, f.map_add], - map_smul' := λ v i c x, by simp only [← sum.update_elim_inr, f.map_smul] }, - map_add' := λ u i x y, ext $ λ v, - by simp only [multilinear_map.coe_mk, add_apply, ← sum.update_elim_inl, f.map_add], - map_smul' := λ u i c x, ext $ λ v, - by simp only [multilinear_map.coe_mk, smul_apply, ← sum.update_elim_inl, f.map_smul] } + map_add' := λ _ v i x y, by + { resetI, letI := classical.dec_eq ι, + simp only [← sum.update_elim_inr, f.map_add] }, + map_smul' := λ _ v i c x, by + { resetI, letI := classical.dec_eq ι, + simp only [← sum.update_elim_inr, f.map_smul] } }, + map_add' := λ _ u i x y, ext $ λ v, by + { resetI, letI := classical.dec_eq ι', + simp only [multilinear_map.coe_mk, add_apply, ← sum.update_elim_inl, f.map_add] }, + map_smul' := λ _ u i c x, ext $ λ v, by + { resetI, letI := classical.dec_eq ι', + simp only [multilinear_map.coe_mk, smul_apply, ← sum.update_elim_inl, f.map_smul] } } @[simp] lemma curry_sum_apply (f : multilinear_map R (λ x : ι ⊕ ι', M') M₂) (u : ι → M') (v : ι' → M') : @@ -1203,12 +1252,20 @@ on `Π i : ι', M'` defines a multilinear map on `Π i : ι ⊕ ι', M'`. -/ def uncurry_sum (f : multilinear_map R (λ x : ι, M') (multilinear_map R (λ x : ι', M') M₂)) : multilinear_map R (λ x : ι ⊕ ι', M') M₂ := { to_fun := λ u, f (u ∘ sum.inl) (u ∘ sum.inr), - map_add' := λ u i x y, by cases i; - simp only [multilinear_map.map_add, add_apply, sum.update_inl_comp_inl, sum.update_inl_comp_inr, - sum.update_inr_comp_inl, sum.update_inr_comp_inr], - map_smul' := λ u i c x, by cases i; + map_add' := λ _ u i x y, by + { resetI, + letI := (@sum.inl_injective ι ι').decidable_eq, + letI := (@sum.inr_injective ι ι').decidable_eq, + cases i; + simp only [multilinear_map.map_add, add_apply, sum.update_inl_comp_inl, + sum.update_inl_comp_inr, sum.update_inr_comp_inl, sum.update_inr_comp_inr] }, + map_smul' := λ _ u i c x, by + { resetI, + letI := (@sum.inl_injective ι ι').decidable_eq, + letI := (@sum.inr_injective ι ι').decidable_eq, + cases i; simp only [multilinear_map.map_smul, smul_apply, sum.update_inl_comp_inl, - sum.update_inl_comp_inr, sum.update_inr_comp_inl, sum.update_inr_comp_inr] } + sum.update_inl_comp_inr, sum.update_inr_comp_inl, sum.update_inr_comp_inr] } } @[simp] lemma uncurry_sum_aux_apply (f : multilinear_map R (λ x : ι, M') (multilinear_map R (λ x : ι', M') M₂)) (u : ι ⊕ ι' → M') : @@ -1312,7 +1369,8 @@ def map [nonempty ι] (f : multilinear_map R M₁ M₂) (p : Π i, submodule R ( sub_mul_action R M₂ := { carrier := f '' { v | ∀ i, v i ∈ p i}, smul_mem' := λ c _ ⟨x, hx, hf⟩, let ⟨i⟩ := ‹nonempty ι› in by - { refine ⟨update x i (c • x i), λ j, if hij : j = i then _ else _, hf ▸ _⟩, + { letI := classical.dec_eq ι, + refine ⟨update x i (c • x i), λ j, if hij : j = i then _ else _, hf ▸ _⟩, { rw [hij, update_same], exact (p i).smul_mem _ (hx i) }, { rw [update_noteq hij], exact hx j }, { rw [f.map_smul, update_eq_self] } } } diff --git a/src/linear_algebra/multilinear/basis.lean b/src/linear_algebra/multilinear/basis.lean index 1b1a06cac5a33..86a1f14dd9d4b 100644 --- a/src/linear_algebra/multilinear/basis.lean +++ b/src/linear_algebra/multilinear/basis.lean @@ -51,7 +51,7 @@ end are basis vectors. Unlike `basis.ext_multilinear_fin`, this only uses a single basis; a dependently-typed version would still be true, but the proof would need a dependently-typed version of `dom_dom_congr`. -/ -lemma basis.ext_multilinear [decidable_eq ι] [finite ι] {f g : multilinear_map R (λ i : ι, M₂) M₃} +lemma basis.ext_multilinear [finite ι] {f g : multilinear_map R (λ i : ι, M₂) M₃} {ι₁ : Type*} (e : basis ι₁ R M₂) (h : ∀ v : ι → ι₁, f (λ i, e (v i)) = g (λ i, e (v i))) : f = g := by { casesI nonempty_fintype ι, exact (dom_dom_congr_eq_iff (fintype.equiv_fin ι) f g).mp diff --git a/src/linear_algebra/multilinear/finite_dimensional.lean b/src/linear_algebra/multilinear/finite_dimensional.lean index fc22a203d60b8..eba4f84553061 100644 --- a/src/linear_algebra/multilinear/finite_dimensional.lean +++ b/src/linear_algebra/multilinear/finite_dimensional.lean @@ -21,7 +21,7 @@ there. namespace multilinear_map variables {ι R M₂ : Type*} {M₁ : ι → Type*} -variables [decidable_eq ι] [finite ι] +variables [finite ι] variables [comm_ring R] [add_comm_group M₂] [module R M₂] variables [Π i, add_comm_group (M₁ i)] [Π i, module R (M₁ i)] variables [module.finite R M₂] [module.free R M₂] diff --git a/src/linear_algebra/multilinear/tensor_product.lean b/src/linear_algebra/multilinear/tensor_product.lean index 184a514b7acc2..07a7bd39f336e 100644 --- a/src/linear_algebra/multilinear/tensor_product.lean +++ b/src/linear_algebra/multilinear/tensor_product.lean @@ -18,7 +18,6 @@ open_locale tensor_product variables {R ι₁ ι₂ ι₃ ι₄ : Type*} variables [comm_semiring R] -variables [decidable_eq ι₁] [decidable_eq ι₂][decidable_eq ι₃] [decidable_eq ι₄] variables {N₁ : Type*} [add_comm_monoid N₁] [module R N₁] variables {N₂ : Type*} [add_comm_monoid N₂] [module R N₂] variables {N : Type*} [add_comm_monoid N] [module R N] @@ -41,8 +40,16 @@ def dom_coprod (a : multilinear_map R (λ _ : ι₁, N) N₁) (b : multilinear_map R (λ _ : ι₂, N) N₂) : multilinear_map R (λ _ : ι₁ ⊕ ι₂, N) (N₁ ⊗[R] N₂) := { to_fun := λ v, a (λ i, v (sum.inl i)) ⊗ₜ b (λ i, v (sum.inr i)), - map_add' := λ v i p q, by cases i; simp [tensor_product.add_tmul, tensor_product.tmul_add], - map_smul' := λ v i c p, by cases i; simp [tensor_product.smul_tmul', tensor_product.tmul_smul] } + map_add' := λ _ v i p q, by + { resetI, + letI := (@sum.inl_injective ι₁ ι₂).decidable_eq, + letI := (@sum.inr_injective ι₁ ι₂).decidable_eq, + cases i; simp [tensor_product.add_tmul, tensor_product.tmul_add] }, + map_smul' := λ _ v i c p, by + { resetI, + letI := (@sum.inl_injective ι₁ ι₂).decidable_eq, + letI := (@sum.inr_injective ι₁ ι₂).decidable_eq, + cases i; simp [tensor_product.smul_tmul', tensor_product.tmul_smul] } } /-- A more bundled version of `multilinear_map.dom_coprod` that maps `((ι₁ → N) → N₁) ⊗ ((ι₂ → N) → N₂)` to `(ι₁ ⊕ ι₂ → N) → N₁ ⊗ N₂`. -/ diff --git a/src/linear_algebra/orientation.lean b/src/linear_algebra/orientation.lean index 3a940f78fb7b9..e706b7ccac3be 100644 --- a/src/linear_algebra/orientation.lean +++ b/src/linear_algebra/orientation.lean @@ -40,7 +40,7 @@ section ordered_comm_semiring variables (R : Type*) [strict_ordered_comm_semiring R] variables (M : Type*) [add_comm_monoid M] [module R M] variables {N : Type*} [add_comm_monoid N] [module R N] -variables (ι : Type*) [decidable_eq ι] +variables (ι : Type*) /-- An orientation of a module, intended to be used when `ι` is a `fintype` with the same cardinality as a basis. -/ @@ -99,14 +99,14 @@ section ordered_comm_ring variables {R : Type*} [strict_ordered_comm_ring R] variables {M N : Type*} [add_comm_group M] [add_comm_group N] [module R M] [module R N] -@[simp] protected lemma orientation.map_neg {ι : Type*} [decidable_eq ι] (f : M ≃ₗ[R] N) +@[simp] protected lemma orientation.map_neg {ι : Type*} (f : M ≃ₗ[R] N) (x : orientation R M ι) : orientation.map ι f (-x) = - orientation.map ι f x := module.ray.map_neg _ x namespace basis -variables {ι : Type*} [decidable_eq ι] +variables {ι : Type*} /-- The value of `orientation.map` when the index type has the cardinality of a basis, in terms of `f.det`. -/ @@ -114,6 +114,7 @@ lemma map_orientation_eq_det_inv_smul [finite ι] (e : basis ι R M) (x : orientation R M ι) (f : M ≃ₗ[R] M) : orientation.map ι f x = (f.det)⁻¹ • x := begin casesI nonempty_fintype ι, + letI := classical.dec_eq ι, induction x using module.ray.ind with g hg, rw [orientation.map_apply, smul_ray_of_ne_zero, ray_eq_iff, units.smul_def, (g.comp_linear_map ↑f.symm).eq_smul_basis_det e, g.eq_smul_basis_det e, @@ -121,7 +122,7 @@ begin basis.det_self, mul_one, smul_eq_mul, mul_comm, mul_smul, linear_equiv.coe_inv_det], end -variables [fintype ι] +variables [fintype ι] [decidable_eq ι] /-- The orientation given by a basis. -/ protected def orientation [nontrivial R] (e : basis ι R M) : orientation R M ι := @@ -157,7 +158,7 @@ section linear_ordered_comm_ring variables {R : Type*} [linear_ordered_comm_ring R] variables {M : Type*} [add_comm_group M] [module R M] -variables {ι : Type*} [decidable_eq ι] +variables {ι : Type*} namespace orientation @@ -186,7 +187,7 @@ end orientation namespace basis -variables [fintype ι] +variables [fintype ι] [decidable_eq ι] /-- The orientations given by two bases are equal if and only if the determinant of one basis with respect to the other is positive. -/ @@ -296,7 +297,7 @@ section linear_ordered_field variables {R : Type*} [linear_ordered_field R] variables {M : Type*} [add_comm_group M] [module R M] -variables {ι : Type*} [decidable_eq ι] +variables {ι : Type*} namespace orientation @@ -312,6 +313,7 @@ lemma eq_or_eq_neg (x₁ x₂ : orientation R M ι) (h : fintype.card ι = finra x₁ = x₂ ∨ x₁ = -x₂ := begin have e := (fin_basis R M).reindex (fintype.equiv_fin_of_card_eq h).symm, + letI := classical.dec_eq ι, rcases e.orientation_eq_or_eq_neg x₁ with h₁|h₁; rcases e.orientation_eq_or_eq_neg x₂ with h₂|h₂; simp [h₁, h₂] @@ -379,12 +381,13 @@ include _i /-- If the index type has cardinality equal to the finite dimension, a basis with the given orientation. -/ -def some_basis [nonempty ι] (x : orientation R M ι) (h : fintype.card ι = finrank R M) : +def some_basis [nonempty ι] [decidable_eq ι] (x : orientation R M ι) + (h : fintype.card ι = finrank R M) : basis ι R M := ((fin_basis R M).reindex (fintype.equiv_fin_of_card_eq h).symm).adjust_to_orientation x /-- `some_basis` gives a basis with the required orientation. -/ -@[simp] lemma some_basis_orientation [nonempty ι] (x : orientation R M ι) +@[simp] lemma some_basis_orientation [nonempty ι] [decidable_eq ι] (x : orientation R M ι) (h : fintype.card ι = finrank R M) : (x.some_basis h).orientation = x := basis.orientation_adjust_to_orientation _ _ diff --git a/src/linear_algebra/pi_tensor_product.lean b/src/linear_algebra/pi_tensor_product.lean index d50bc7c91cf76..9e0ef6666844b 100644 --- a/src/linear_algebra/pi_tensor_product.lean +++ b/src/linear_algebra/pi_tensor_product.lean @@ -42,6 +42,10 @@ binary tensor product in `linear_algebra/tensor_product.lean`. * We have not restricted the index type `ι` to be a `fintype`, as nothing we do here strictly requires it. However, problems may arise in the case where `ι` is infinite; use at your own caution. +* Instead of requiring `decidable_eq ι` as an argument to `pi_tensor_product` itself, we include it + as an argument in the constructors of the relation. A decidability isntance still has to come + from somewhere due to the use of `function.update`, but this hides it from the downstream user. + See the implementation notes for `multilinear_map` for an extended discussion of this choice. ## TODO @@ -60,7 +64,7 @@ open function section semiring -variables {ι ι₂ ι₃ : Type*} [decidable_eq ι] [decidable_eq ι₂] [decidable_eq ι₃] +variables {ι ι₂ ι₃ : Type*} variables {R : Type*} [comm_semiring R] variables {R₁ R₂ : Type*} variables {s : ι → Type*} [∀ i, add_comm_monoid (s i)] [∀ i, module R (s i)] @@ -77,13 +81,13 @@ the tensor product. -/ inductive eqv : free_add_monoid (R × Π i, s i) → free_add_monoid (R × Π i, s i) → Prop | of_zero : ∀ (r : R) (f : Π i, s i) (i : ι) (hf : f i = 0), eqv (free_add_monoid.of (r, f)) 0 | of_zero_scalar : ∀ (f : Π i, s i), eqv (free_add_monoid.of (0, f)) 0 -| of_add : ∀ (r : R) (f : Π i, s i) (i : ι) (m₁ m₂ : s i), eqv +| of_add : ∀ (inst : decidable_eq ι) (r : R) (f : Π i, s i) (i : ι) (m₁ m₂ : s i), eqv (free_add_monoid.of (r, update f i m₁) + free_add_monoid.of (r, update f i m₂)) (free_add_monoid.of (r, update f i (m₁ + m₂))) | of_add_scalar : ∀ (r r' : R) (f : Π i, s i), eqv (free_add_monoid.of (r, f) + free_add_monoid.of (r', f)) (free_add_monoid.of (r + r', f)) -| of_smul : ∀ (r : R) (f : Π i, s i) (i : ι) (r' : R), eqv +| of_smul : ∀ (inst : decidable_eq ι) (r : R) (f : Π i, s i) (i : ι) (r' : R), eqv (free_add_monoid.of (r, update f i (r' • (f i)))) (free_add_monoid.of (r' * r, f)) | add_comm : ∀ x y, eqv (x + y) (y + x) @@ -131,20 +135,20 @@ quotient.sound' $ add_con_gen.rel.of _ _ $ eqv.of_zero_scalar _ lemma zero_tprod_coeff' (z : R) (f : Π i, s i) (i : ι) (hf: f i = 0) : tprod_coeff R z f = 0 := quotient.sound' $ add_con_gen.rel.of _ _ $ eqv.of_zero _ _ i hf -lemma add_tprod_coeff (z : R) (f : Π i, s i) (i : ι) (m₁ m₂ : s i) : +lemma add_tprod_coeff [decidable_eq ι] (z : R) (f : Π i, s i) (i : ι) (m₁ m₂ : s i) : tprod_coeff R z (update f i m₁) + tprod_coeff R z (update f i m₂) = tprod_coeff R z (update f i (m₁ + m₂)) := -quotient.sound' $ add_con_gen.rel.of _ _ (eqv.of_add z f i m₁ m₂) +quotient.sound' $ add_con_gen.rel.of _ _ (eqv.of_add _ z f i m₁ m₂) lemma add_tprod_coeff' (z₁ z₂ : R) (f : Π i, s i) : tprod_coeff R z₁ f + tprod_coeff R z₂ f = tprod_coeff R (z₁ + z₂) f := quotient.sound' $ add_con_gen.rel.of _ _ (eqv.of_add_scalar z₁ z₂ f) -lemma smul_tprod_coeff_aux (z : R) (f : Π i, s i) (i : ι) (r : R) : +lemma smul_tprod_coeff_aux [decidable_eq ι] (z : R) (f : Π i, s i) (i : ι) (r : R) : tprod_coeff R z (update f i (r • f i)) = tprod_coeff R (r * z) f := - quotient.sound' $ add_con_gen.rel.of _ _ $ eqv.of_smul _ _ _ _ + quotient.sound' $ add_con_gen.rel.of _ _ $ eqv.of_smul _ _ _ _ _ -lemma smul_tprod_coeff (z : R) (f : Π i, s i) (i : ι) (r : R₁) +lemma smul_tprod_coeff [decidable_eq ι] (z : R) (f : Π i, s i) (i : ι) (r : R₁) [has_smul R₁ R] [is_scalar_tower R₁ R R] [has_smul R₁ (s i)] [is_scalar_tower R₁ R (s i)] : tprod_coeff R z (update f i (r • f i)) = tprod_coeff R (r • z) f := begin @@ -159,11 +163,11 @@ end def lift_add_hom (φ : (R × Π i, s i) → F) (C0 : ∀ (r : R) (f : Π i, s i) (i : ι) (hf : f i = 0), φ (r, f) = 0) (C0' : ∀ (f : Π i, s i), φ (0, f) = 0) - (C_add : ∀ (r : R) (f : Π i, s i) (i : ι) (m₁ m₂ : s i), + (C_add : ∀ [decidable_eq ι] (r : R) (f : Π i, s i) (i : ι) (m₁ m₂ : s i), by exactI φ (r, update f i m₁) + φ (r, update f i m₂) = φ (r, update f i (m₁ + m₂))) (C_add_scalar : ∀ (r r' : R) (f : Π i, s i), φ (r , f) + φ (r', f) = φ (r + r', f)) - (C_smul : ∀ (r : R) (f : Π i, s i) (i : ι) (r' : R), + (C_smul : ∀ [decidable_eq ι] (r : R) (f : Π i, s i) (i : ι) (r' : R), by exactI φ (r, update f i (r' • (f i))) = φ (r' * r, f)) : (⨂[R] i, s i) →+ F := (add_con_gen (pi_tensor_product.eqv R s)).lift (free_add_monoid.lift φ) $ add_con.add_con_gen_le $ @@ -172,12 +176,12 @@ def lift_add_hom (φ : (R × Π i, s i) → F) by simp [free_add_monoid.lift_eval_of, C0 r' f i hf] | _, _, (eqv.of_zero_scalar f) := (add_con.ker_rel _).2 $ by simp [free_add_monoid.lift_eval_of, C0'] -| _, _, (eqv.of_add z f i m₁ m₂) := (add_con.ker_rel _).2 $ - by simp [free_add_monoid.lift_eval_of, C_add] +| _, _, (eqv.of_add inst z f i m₁ m₂) := (add_con.ker_rel _).2 $ + by simp [free_add_monoid.lift_eval_of, @C_add inst] | _, _, (eqv.of_add_scalar z₁ z₂ f) := (add_con.ker_rel _).2 $ by simp [free_add_monoid.lift_eval_of, C_add_scalar] -| _, _, (eqv.of_smul z f i r') := (add_con.ker_rel _).2 $ - by simp [free_add_monoid.lift_eval_of, C_smul] +| _, _, (eqv.of_smul inst z f i r') := (add_con.ker_rel _).2 $ + by simp [free_add_monoid.lift_eval_of, @C_smul inst] | _, _, (eqv.add_comm x y) := (add_con.ker_rel _).2 $ by simp_rw [add_monoid_hom.map_add, add_comm] end @@ -269,9 +273,9 @@ variables (R) /-- The canonical `multilinear_map R s (⨂[R] i, s i)`. -/ def tprod : multilinear_map R s (⨂[R] i, s i) := { to_fun := tprod_coeff R 1, - map_add' := λ f i x y, (add_tprod_coeff (1 : R) f i x y).symm, - map_smul' := λ f i r x, - by simp_rw [smul_tprod_coeff', ←smul_tprod_coeff (1 : R) _ i, update_idem, update_same] } + map_add' := λ _ f i x y, by exactI (add_tprod_coeff (1 : R) f i x y).symm, + map_smul' := λ _ f i r x, by + resetI; simp_rw [smul_tprod_coeff', ←smul_tprod_coeff (1 : R) _ i, update_idem, update_same] } variables {R} @@ -323,9 +327,9 @@ def lift_aux (φ : multilinear_map R s E) : (⨂[R] i, s i) →+ E := lift_add_hom (λ (p : R × Π i, s i), p.1 • (φ p.2)) (λ z f i hf, by rw [map_coord_zero φ i hf, smul_zero]) (λ f, by rw [zero_smul]) - (λ z f i m₁ m₂, by rw [←smul_add, φ.map_add]) + (λ _ z f i m₁ m₂, by { resetI, rw [←smul_add, φ.map_add] }) (λ z₁ z₂ f, by rw [←add_smul]) - (λ z f i r, by simp [φ.map_smul, smul_smul, mul_comm]) + (λ _ z f i r, by { resetI, simp [φ.map_smul, smul_smul, mul_comm] }) lemma lift_aux_tprod (φ : multilinear_map R s E) (f : Π i, s i) : lift_aux φ (tprod R f) = φ f := by simp only [lift_aux, lift_add_hom, tprod, multilinear_map.coe_mk, tprod_coeff, @@ -558,7 +562,7 @@ namespace pi_tensor_product open pi_tensor_product open_locale tensor_product -variables {ι : Type*} [decidable_eq ι] {R : Type*} [comm_ring R] +variables {ι : Type*} {R : Type*} [comm_ring R] variables {s : ι → Type*} [∀ i, add_comm_group (s i)] [∀ i, module R (s i)] /- Unlike for the binary tensor product, we require `R` to be a `comm_ring` here, otherwise diff --git a/src/linear_algebra/tensor_power.lean b/src/linear_algebra/tensor_power.lean index 828dda9d634fb..a17d3b28dab14 100644 --- a/src/linear_algebra/tensor_power.lean +++ b/src/linear_algebra/tensor_power.lean @@ -46,8 +46,7 @@ namespace pi_tensor_product /-- Two dependent pairs of tensor products are equal if their index is equal and the contents are equal after a canonical reindexing. -/ @[ext] -lemma graded_monoid_eq_of_reindex_cast {ιι : Type*} {ι : ιι → Type*} - [dι : Π ii, decidable_eq (ι ii)] : +lemma graded_monoid_eq_of_reindex_cast {ιι : Type*} {ι : ιι → Type*} : ∀ {a b : graded_monoid (λ ii, ⨂[R] i : ι ii, M)} (h : a.fst = b.fst), reindex R M (equiv.cast $ congr_arg ι h) a.snd = b.snd → a = b | ⟨ai, a⟩ ⟨bi, b⟩ := λ (hi : ai = bi) (h : reindex R M _ a = b), diff --git a/src/topology/algebra/module/multilinear.lean b/src/topology/algebra/module/multilinear.lean index 5cc8f9997f417..0559849b68edd 100644 --- a/src/topology/algebra/module/multilinear.lean +++ b/src/topology/algebra/module/multilinear.lean @@ -36,14 +36,14 @@ open_locale big_operators universes u v w w₁ w₁' w₂ w₃ w₄ variables {R : Type u} {ι : Type v} {n : ℕ} {M : fin n.succ → Type w} {M₁ : ι → Type w₁} - {M₁' : ι → Type w₁'} {M₂ : Type w₂} {M₃ : Type w₃} {M₄ : Type w₄} [decidable_eq ι] + {M₁' : ι → Type w₁'} {M₂ : Type w₂} {M₃ : Type w₃} {M₄ : Type w₄} /-- Continuous multilinear maps over the ring `R`, from `Πi, M₁ i` to `M₂` where `M₁ i` and `M₂` are modules over `R` with a topological structure. In applications, there will be compatibility conditions between the algebraic and the topological structures, but this is not needed for the definition. -/ structure continuous_multilinear_map (R : Type u) {ι : Type v} (M₁ : ι → Type w₁) (M₂ : Type w₂) - [decidable_eq ι] [semiring R] [∀i, add_comm_monoid (M₁ i)] [add_comm_monoid M₂] + [semiring R] [∀i, add_comm_monoid (M₁ i)] [add_comm_monoid M₂] [∀i, module R (M₁ i)] [module R M₂] [∀i, topological_space (M₁ i)] [topological_space M₂] extends multilinear_map R M₁ M₂ := (cont : continuous to_fun) @@ -88,11 +88,11 @@ to_multilinear_map_inj $ multilinear_map.ext H theorem ext_iff {f f' : continuous_multilinear_map R M₁ M₂} : f = f' ↔ ∀ x, f x = f' x := by rw [← to_multilinear_map_inj.eq_iff, multilinear_map.ext_iff]; refl -@[simp] lemma map_add (m : Πi, M₁ i) (i : ι) (x y : M₁ i) : +@[simp] lemma map_add [decidable_eq ι] (m : Πi, M₁ i) (i : ι) (x y : M₁ i) : f (update m i (x + y)) = f (update m i x) + f (update m i y) := f.map_add' m i x y -@[simp] lemma map_smul (m : Πi, M₁ i) (i : ι) (c : R) (x : M₁ i) : +@[simp] lemma map_smul [decidable_eq ι] (m : Πi, M₁ i) (i : ι) (c : R) (x : M₁ i) : f (update m i (c • x)) = c • f (update m i x) := f.map_smul' m i c x @@ -174,7 +174,7 @@ end has_continuous_add /-- If `f` is a continuous multilinear map, then `f.to_continuous_linear_map m i` is the continuous linear map obtained by fixing all coordinates but `i` equal to those of `m`, and varying the `i`-th coordinate. -/ -def to_continuous_linear_map (m : Πi, M₁ i) (i : ι) : M₁ i →L[R] M₂ := +def to_continuous_linear_map [decidable_eq ι] (m : Πi, M₁ i) (i : ι) : M₁ i →L[R] M₂ := { cont := f.cont.comp (continuous_const.update i continuous_id), .. f.to_multilinear_map.to_linear_map m i } @@ -282,13 +282,13 @@ lemma cons_smul f (cons (c • x) m) = c • f (cons x m) := f.to_multilinear_map.cons_smul m c x -lemma map_piecewise_add (m m' : Πi, M₁ i) (t : finset ι) : +lemma map_piecewise_add [decidable_eq ι] (m m' : Πi, M₁ i) (t : finset ι) : f (t.piecewise (m + m') m') = ∑ s in t.powerset, f (s.piecewise m m') := f.to_multilinear_map.map_piecewise_add _ _ _ /-- Additivity of a continuous multilinear map along all coordinates at the same time, writing `f (m + m')` as the sum of `f (s.piecewise m m')` over all sets `s`. -/ -lemma map_add_univ [fintype ι] (m m' : Πi, M₁ i) : +lemma map_add_univ [decidable_eq ι] [fintype ι] (m m' : Πi, M₁ i) : f (m + m') = ∑ s : finset ι, f (s.piecewise m m') := f.to_multilinear_map.map_add_univ _ _ @@ -303,14 +303,14 @@ sum of `f (g₁ (r 1), ..., gₙ (r n))` where `r` ranges over all functions wit `r n ∈ Aₙ`. This follows from multilinearity by expanding successively with respect to each coordinate. -/ -lemma map_sum_finset : +lemma map_sum_finset [decidable_eq ι] : f (λ i, ∑ j in A i, g i j) = ∑ r in pi_finset A, f (λ i, g i (r i)) := f.to_multilinear_map.map_sum_finset _ _ /-- If `f` is continuous multilinear, then `f (Σ_{j₁} g₁ j₁, ..., Σ_{jₙ} gₙ jₙ)` is the sum of `f (g₁ (r 1), ..., gₙ (r n))` where `r` ranges over all functions `r`. This follows from multilinearity by expanding successively with respect to each coordinate. -/ -lemma map_sum [∀ i, fintype (α i)] : +lemma map_sum [decidable_eq ι] [∀ i, fintype (α i)] : f (λ i, ∑ j, g i j) = ∑ r : Π i, α i, f (λ i, g i (r i)) := f.to_multilinear_map.map_sum _ @@ -341,7 +341,7 @@ variables [ring R] [∀i, add_comm_group (M₁ i)] [add_comm_group M₂] [∀i, module R (M₁ i)] [module R M₂] [∀i, topological_space (M₁ i)] [topological_space M₂] (f f' : continuous_multilinear_map R M₁ M₂) -@[simp] lemma map_sub (m : Πi, M₁ i) (i : ι) (x y : M₁ i) : +@[simp] lemma map_sub [decidable_eq ι] (m : Πi, M₁ i) (i : ι) (x y : M₁ i) : f (update m i (x - y)) = f (update m i x) - f (update m i y) := f.to_multilinear_map.map_sub _ _ _ _ @@ -374,7 +374,7 @@ variables [comm_semiring R] [∀i, topological_space (M₁ i)] [topological_space M₂] (f : continuous_multilinear_map R M₁ M₂) -lemma map_piecewise_smul (c : ι → R) (m : Πi, M₁ i) (s : finset ι) : +lemma map_piecewise_smul [decidable_eq ι] (c : ι → R) (m : Πi, M₁ i) (s : finset ι) : f (s.piecewise (λ i, c i • m i) m) = (∏ i in s, c i) • f m := f.to_multilinear_map.map_piecewise_smul _ _ _ From 3e068ece210655b7b9a9477c3aff38a492400aa1 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 30 Mar 2023 21:32:22 +0000 Subject: [PATCH 0741/1291] refactor(data/matrix/basic): work around leanprover/lean4#2042 (#18696) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This adjust definitions such that everything is well-behaved in the case that things are unfolded. For each such definition, a lemma is added that replaces the equation lemma. Before this PR, we used ```lean def transpose (M : matrix m n α) : matrix n m α | x y := M y x ``` which has the nice behavior (in Lean 3 only) of `rw transpose` only unfolding the definition when it is of the applied form `transpose M i j`. If `dunfold transpose` is used then it becomes the undesirable `λ x y, M y x` in both Lean versions. After this PR, we use ```lean def transpose (M : matrix m n α) : matrix n m α := of $ λ x y, M y x -- TODO: set as an equation lemma for `transpose`, see mathlib4#3024 @[simp] lemma transpose_apply (M : matrix m n α) (i j) : transpose M i j = M j i := rfl ``` This no longer has the nice `rw` behavior, but we can't have that in Lean4 anyway (leanprover/lean4#2042). It also makes `dunfold` insert the `of`, which is better for type-safety. This affects * `matrix.transpose` * `matrix.row` * `matrix.col` * `matrix.diagonal` * `matrix.vec_mul_vec` * `matrix.block_diagonal` * `matrix.block_diagonal'` * `matrix.hadamard` * `matrix.kronecker_map` * `pequiv.to_matrix` * `matrix.circulant` * `matrix.mv_polynomial_X` * `algebra.trace_matrix` * `algebra.embeddings_matrix` While this just adds `_apply` noise in Lean 3, it is necessary when porting to Lean 4 as there the equation lemma is not generated in the way that we want. This is hopefully exhaustive; it was found by looking for lines ending in `matrix .*` followed by a `|` line --- src/algebra/lie/classical.lean | 2 +- .../simple_graph/adj_matrix.lean | 5 +- src/data/matrix/basic.lean | 66 ++++++++++++------- src/data/matrix/block.lean | 50 +++++++++----- src/data/matrix/hadamard.lean | 9 ++- src/data/matrix/kronecker.lean | 15 +++-- src/data/matrix/notation.lean | 2 +- src/data/matrix/pequiv.lean | 15 +++-- src/linear_algebra/matrix/circulant.lean | 11 ++-- src/linear_algebra/matrix/mv_polynomial.lean | 9 ++- src/linear_algebra/matrix/symmetric.lean | 4 +- src/linear_algebra/matrix/to_lin.lean | 2 +- .../cyclotomic/discriminant.lean | 2 +- src/ring_theory/discriminant.lean | 6 +- src/ring_theory/trace.lean | 25 ++++--- src/topology/instances/matrix.lean | 2 +- 16 files changed, 145 insertions(+), 80 deletions(-) diff --git a/src/algebra/lie/classical.lean b/src/algebra/lie/classical.lean index 62b341b45b446..7513bfe024462 100644 --- a/src/algebra/lie/classical.lean +++ b/src/algebra/lie/classical.lean @@ -328,7 +328,7 @@ begin ext i j, rcases i with ⟨⟨i₁ | i₂⟩ | i₃⟩; rcases j with ⟨⟨j₁ | j₂⟩ | j₃⟩; - simp only [indefinite_diagonal, matrix.diagonal, equiv.sum_assoc_apply_inl_inl, + simp only [indefinite_diagonal, matrix.diagonal_apply, equiv.sum_assoc_apply_inl_inl, matrix.reindex_lie_equiv_apply, matrix.submatrix_apply, equiv.symm_symm, matrix.reindex_apply, sum.elim_inl, if_true, eq_self_iff_true, matrix.one_apply_eq, matrix.from_blocks_apply₁₁, dmatrix.zero_apply, equiv.sum_assoc_apply_inl_inr, if_false, matrix.from_blocks_apply₁₂, diff --git a/src/combinatorics/simple_graph/adj_matrix.lean b/src/combinatorics/simple_graph/adj_matrix.lean index 4f851368c63e7..fbb570bf19eb3 100644 --- a/src/combinatorics/simple_graph/adj_matrix.lean +++ b/src/combinatorics/simple_graph/adj_matrix.lean @@ -148,11 +148,12 @@ variables (α) /-- `adj_matrix G α` is the matrix `A` such that `A i j = (1 : α)` if `i` and `j` are adjacent in the simple graph `G`, and otherwise `A i j = 0`. -/ -def adj_matrix [has_zero α] [has_one α] : matrix V V α -| i j := if (G.adj i j) then 1 else 0 +def adj_matrix [has_zero α] [has_one α] : matrix V V α := +of $ λ i j, if (G.adj i j) then (1 : α) else 0 variable {α} +-- TODO: set as an equation lemma for `adj_matrix`, see mathlib4#3024 @[simp] lemma adj_matrix_apply (v w : V) [has_zero α] [has_one α] : G.adj_matrix α v w = if (G.adj v w) then 1 else 0 := rfl diff --git a/src/data/matrix/basic.lean b/src/data/matrix/basic.lean index 85e12b44c0ac0..76c5cb08b2861 100644 --- a/src/data/matrix/basic.lean +++ b/src/data/matrix/basic.lean @@ -76,10 +76,14 @@ The two sides of the equivalence are definitionally equal types. We want to use to distinguish the types because `matrix` has different instances to pi types (such as `pi.has_mul`, which performs elementwise multiplication, vs `matrix.has_mul`). -If you are defining a matrix, in terms of its entries, either use `of (λ i j, _)`, or use pattern -matching in a definition as `| i j := _` (which can only be unfolded when fully-applied). The -purpose of this approach is to ensure that terms of the form `(λ i j, _) * (λ i j, _)` do not +If you are defining a matrix, in terms of its entries, use `of (λ i j, _)`. The +purpose of this approach is to ensure that terms of th +e form `(λ i j, _) * (λ i j, _)` do not appear, as the type of `*` can be misleading. + +Porting note: In Lean 3, it is also safe to use pattern matching in a definition as `| i j := _`, +which can only be unfolded when fully-applied. leanprover/lean4#2042 means this does not +(currently) work in Lean 4. -/ def of : (m → n → α) ≃ matrix m n α := equiv.refl _ @[simp] lemma of_apply (f : m → n → α) (i j) : of f i j = f i j := rfl @@ -118,8 +122,12 @@ lemma map_injective {f : α → β} (hf : function.injective 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 +def transpose (M : matrix m n α) : matrix n m α := +of $ λ x y, M y x + +-- TODO: set as an equation lemma for `transpose`, see mathlib4#3024 +@[simp] lemma transpose_apply (M : matrix m n α) (i j) : + transpose M i j = M j i := rfl localized "postfix (name := matrix.transpose) `ᵀ`:1500 := matrix.transpose" in matrix @@ -130,12 +138,19 @@ M.transpose.map star localized "postfix (name := matrix.conj_transpose) `ᴴ`:1500 := matrix.conj_transpose" in matrix /-- `matrix.col u` is the column matrix whose entries are given by `u`. -/ -def col (w : m → α) : matrix m unit α -| x y := w x +def col (w : m → α) : matrix m unit α := +of $ λ x y, w x + +-- TODO: set as an equation lemma for `col`, see mathlib4#3024 +@[simp] lemma col_apply (w : m → α) (i j) : + col w i j = w i := rfl /-- `matrix.row u` is the row matrix whose entries are given by `u`. -/ -def row (v : n → α) : matrix unit n α -| x y := v y +def row (v : n → α) : matrix unit n α := +of $ λ x y, v y + +-- TODO: set as an equation lemma for `row`, see mathlib4#3024 +@[simp] lemma row_apply (v : n → α) (i j) : row v i j = v j := rfl instance [inhabited α] : inhabited (matrix m n α) := pi.inhabited _ instance [has_add α] : has_add (matrix m n α) := pi.has_add @@ -239,8 +254,12 @@ Note that bundled versions exist as: * `matrix.diagonal_ring_hom` * `matrix.diagonal_alg_hom` -/ -def diagonal [has_zero α] (d : n → α) : matrix n n α -| i j := if i = j then d i else 0 +def diagonal [has_zero α] (d : n → α) : matrix n n α := +of $ λ i j, if i = j then d i else 0 + +-- TODO: set as an equation lemma for `diagonal`, see mathlib4#3024 +lemma diagonal_apply [has_zero α] (d : n → α) (i j) : diagonal d i j = if i = j then d i else 0 := +rfl @[simp] theorem diagonal_apply_eq [has_zero α] (d : n → α) (i : n) : (diagonal d) i i = d i := by simp [diagonal] @@ -302,7 +321,7 @@ variables {n α R} @[simp] lemma diagonal_map [has_zero α] [has_zero β] {f : α → β} (h : f 0 = 0) {d : n → α} : (diagonal d).map f = diagonal (λ m, f (d m)) := -by { ext, simp only [diagonal, map_apply], split_ifs; simp [h], } +by { ext, simp only [diagonal_apply, map_apply], split_ifs; simp [h], } @[simp] lemma diagonal_conj_transpose [add_monoid α] [star_add_monoid α] (v : n → α) : (diagonal v)ᴴ = diagonal (star v) := @@ -1113,8 +1132,13 @@ namespace matrix /-- For two vectors `w` and `v`, `vec_mul_vec w v i j` is defined to be `w i * v j`. Put another way, `vec_mul_vec w v` is exactly `col w ⬝ row v`. -/ -def vec_mul_vec [has_mul α] (w : m → α) (v : n → α) : matrix m n α -| x y := w x * v y +def vec_mul_vec [has_mul α] (w : m → α) (v : n → α) : matrix m n α := +of $ λ x y, w x * v y + +-- TODO: set as an equation lemma for `vec_mul_vec`, see mathlib4#3024 +lemma vec_mul_vec_apply [has_mul α] (w : m → α) (v : n → α) (i j) : + vec_mul_vec w v i j = w i * v j := +rfl lemma vec_mul_vec_eq [has_mul α] [add_comm_monoid α] (w : m → α) (v : n → α) : vec_mul_vec w v = (col w) ⬝ (row v) := @@ -1336,13 +1360,6 @@ section transpose open_locale matrix -/-- - Tell `simp` what the entries are in a transposed matrix. - - Compare with `mul_apply`, `diagonal_apply_eq`, etc. --/ -@[simp] lemma transpose_apply (M : matrix m n α) (i j) : M.transpose j i = M i j := rfl - @[simp] lemma transpose_transpose (M : matrix m n α) : Mᵀᵀ = M := by ext; refl @@ -1353,7 +1370,7 @@ by ext i j; refl @[simp] lemma transpose_one [decidable_eq n] [has_zero α] [has_one α] : (1 : matrix n n α)ᵀ = 1 := begin ext i j, - unfold has_one.one transpose, + rw [transpose_apply, ←diagonal_one], by_cases i = j, { simp only [h, diagonal_apply_eq] }, { simp only [diagonal_apply_ne _ h, diagonal_apply_ne' _ h] } @@ -1435,6 +1452,8 @@ def transpose_ring_equiv [add_comm_monoid α] [comm_semigroup α] [fintype m] : inv_fun := λ M, M.unopᵀ, map_mul' := λ M N, (congr_arg mul_opposite.op (transpose_mul M N)).trans (mul_opposite.op_mul _ _), + left_inv := λ M, transpose_transpose M, + right_inv := λ M, mul_opposite.unop_injective $ transpose_transpose M.unop, ..(transpose_add_equiv m m α).trans mul_opposite.op_add_equiv } variables {m α} @@ -1895,9 +1914,6 @@ by { ext, refl } @[simp] lemma row_smul [has_smul R α] (x : R) (v : m → α) : row (x • v) = x • row v := by { ext, refl } -@[simp] lemma col_apply (v : m → α) (i j) : matrix.col v i j = v i := rfl -@[simp] lemma row_apply (v : m → α) (i j) : matrix.row v i j = v j := rfl - @[simp] lemma transpose_col (v : m → α) : (matrix.col v)ᵀ = matrix.row v := by { ext, refl } @[simp] diff --git a/src/data/matrix/block.lean b/src/data/matrix/block.lean index 911895a970938..d51d3e867abc4 100644 --- a/src/data/matrix/block.lean +++ b/src/data/matrix/block.lean @@ -278,8 +278,13 @@ the diagonal and zero elsewhere. See also `matrix.block_diagonal'` if the matrices may not have the same size everywhere. -/ -def block_diagonal (M : o → matrix m n α) : matrix (m × o) (n × o) α -| ⟨i, k⟩ ⟨j, k'⟩ := if k = k' then M k i j else 0 +def block_diagonal (M : o → matrix m n α) : matrix (m × o) (n × o) α := +of $ (λ ⟨i, k⟩ ⟨j, k'⟩, if k = k' then M k i j else 0 : m × o → n × o → α) + +-- TODO: set as an equation lemma for `block_diagonal`, see mathlib4#3024 +lemma block_diagonal_apply' (M : o → matrix m n α) (i k j k') : + block_diagonal M ⟨i, k⟩ ⟨j, k'⟩ = if k = k' then M k i j else 0 := +rfl lemma block_diagonal_apply (M : o → matrix m n α) (ik jk) : block_diagonal M ik jk = if ik.2 = jk.2 then M ik.2 ik.1 jk.1 else 0 := @@ -328,7 +333,7 @@ by { ext, simp [block_diagonal_apply] } block_diagonal (λ k, diagonal (d k)) = diagonal (λ ik, d ik.2 ik.1) := begin ext ⟨i, k⟩ ⟨j, k'⟩, - simp only [block_diagonal_apply, diagonal, prod.mk.inj_iff, ← ite_and], + simp only [block_diagonal_apply, diagonal_apply, prod.mk.inj_iff, ← ite_and], congr' 1, rw and_comm, end @@ -404,8 +409,12 @@ 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) +def block_diag (M : matrix (m × o) (n × o) α) (k : o) : matrix m n α := +of $ λ i j, M (i, k) (j, k) + +-- TODO: set as an equation lemma for `block_diag`, see mathlib4#3024 +lemma block_diag_apply (M : matrix (m × o) (n × o) α) (k : o) (i j) : + block_diag M k i j = M (i, k) (j, k) := rfl lemma block_diag_map (M : matrix (m × o) (n × o) α) (f : α → β) : block_diag (M.map f) = λ k, (block_diag M k).map f := @@ -431,14 +440,14 @@ rfl 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)], + { rw [block_diag_apply, diagonal_apply_eq, diagonal_apply_eq] }, + { rw [block_diag_apply, 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 _ _ _ _ +funext $ λ k, ext $ λ i j, block_diagonal_apply_eq M i j _ @[simp] lemma block_diag_one [decidable_eq o] [decidable_eq m] [has_one α] : block_diag (1 : matrix (m × o) (m × o) α) = 1 := @@ -486,8 +495,15 @@ variables [has_zero α] [has_zero β] and zero elsewhere. This is the dependently-typed version of `matrix.block_diagonal`. -/ -def block_diagonal' (M : Π i, matrix (m' i) (n' i) α) : matrix (Σ i, m' i) (Σ i, n' i) α -| ⟨k, i⟩ ⟨k', j⟩ := if h : k = k' then M k i (cast (congr_arg n' h.symm) j) else 0 +def block_diagonal' (M : Π i, matrix (m' i) (n' i) α) : matrix (Σ i, m' i) (Σ i, n' i) α := +of $ (λ ⟨k, i⟩ ⟨k', j⟩, if h : k = k' then M k i (cast (congr_arg n' h.symm) j) else 0 : + (Σ i, m' i) → (Σ i, n' i) → α) + +-- TODO: set as an equation lemma for `block_diagonal'`, see mathlib4#3024 +lemma block_diagonal'_apply' (M : Π i, matrix (m' i) (n' i) α) (k i k' j) : + block_diagonal' M ⟨k, i⟩ ⟨k', j⟩ = + if h : k = k' then M k i (cast (congr_arg n' h.symm) j) else 0 := +rfl lemma block_diagonal'_eq_block_diagonal (M : o → matrix m n α) {k k'} (i j) : block_diagonal M (i, k) (j, k') = block_diagonal' M ⟨k, i⟩ ⟨k', j⟩ := @@ -625,8 +641,12 @@ 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⟩ +def block_diag' (M : matrix (Σ i, m' i) (Σ i, n' i) α) (k : o) : matrix (m' k) (n' k) α := +of $ λ i j, M ⟨k, i⟩ ⟨k, j⟩ + +-- TODO: set as an equation lemma for `block_diag'`, see mathlib4#3024 +lemma block_diag'_apply (M : matrix (Σ i, m' i) (Σ i, n' i) α) (k : o) (i j) : + block_diag' M k i j = M ⟨k, i⟩ ⟨k, j⟩ := rfl 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 := @@ -653,14 +673,14 @@ rfl 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)], + { rw [block_diag'_apply, diagonal_apply_eq, diagonal_apply_eq] }, + { rw [block_diag'_apply, 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 _ _ _ _ +funext $ λ k, ext $ λ i j, block_diagonal'_apply_eq M _ _ _ @[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 := diff --git a/src/data/matrix/hadamard.lean b/src/data/matrix/hadamard.lean index b93f500f0de78..a95155afaf89e 100644 --- a/src/data/matrix/hadamard.lean +++ b/src/data/matrix/hadamard.lean @@ -37,10 +37,13 @@ open_locale matrix big_operators /-- `matrix.hadamard` defines the Hadamard product, which is the pointwise product of two matrices of the same size.-/ -@[simp] -def hadamard [has_mul α] (A : matrix m n α) (B : matrix m n α) : matrix m n α -| i j := A i j * B i j +def hadamard [has_mul α] (A : matrix m n α) (B : matrix m n α) : matrix m n α := +of $ λ i j, A i j * B i j +-- TODO: set as an equation lemma for `hadamard`, see mathlib4#3024 +@[simp] +lemma hadamard_apply [has_mul α] (A : matrix m n α) (B : matrix m n α) (i j) : + hadamard A B i j = A i j * B i j := rfl localized "infix (name := matrix.hadamard) ` ⊙ `:100 := matrix.hadamard" in matrix section basic_properties diff --git a/src/data/matrix/kronecker.lean b/src/data/matrix/kronecker.lean index dfe8fa14ac296..9c07d87b561bf 100644 --- a/src/data/matrix/kronecker.lean +++ b/src/data/matrix/kronecker.lean @@ -52,9 +52,14 @@ variables {l m n p : Type*} {q r : Type*} {l' m' n' p' : Type*} section kronecker_map /-- Produce a matrix with `f` applied to every pair of elements from `A` and `B`. -/ -@[simp] def kronecker_map (f : α → β → γ) (A : matrix l m α) (B : matrix n p β) : - matrix (l × n) (m × p) γ -| i j := f (A i.1 j.1) (B i.2 j.2) +def kronecker_map (f : α → β → γ) (A : matrix l m α) (B : matrix n p β) : + matrix (l × n) (m × p) γ := +of $ λ (i : l × n) (j : m × p), f (A i.1 j.1) (B i.2 j.2) + +-- TODO: set as an equation lemma for `kronecker_map`, see mathlib4#3024 +@[simp] +lemma kronecker_map_apply (f : α → β → γ) (A : matrix l m α) (B : matrix n p β) (i j) : + kronecker_map f A B i j = f (A i.1 j.1) (B i.2 j.2) := rfl lemma kronecker_map_transpose (f : α → β → γ) (A : matrix l m α) (B : matrix n p β) : @@ -199,7 +204,7 @@ lemma kronecker_map_bilinear_mul_mul [comm_semiring R] begin ext ⟨i, i'⟩ ⟨j, j'⟩, simp only [kronecker_map_bilinear_apply_apply, mul_apply, ← finset.univ_product_univ, - finset.sum_product, kronecker_map], + finset.sum_product, kronecker_map_apply], simp_rw [f.map_sum, linear_map.sum_apply, linear_map.map_sum, h_comm], end @@ -212,7 +217,7 @@ lemma trace_kronecker_map_bilinear [comm_semiring R] (f : α →ₗ[R] β →ₗ[R] γ) (A : matrix m m α) (B : matrix n n β) : trace (kronecker_map_bilinear f A B) = f (trace A) (trace B) := by simp_rw [matrix.trace, matrix.diag, kronecker_map_bilinear_apply_apply, - linear_map.map_sum₂, map_sum, ←finset.univ_product_univ, finset.sum_product, kronecker_map] + linear_map.map_sum₂, map_sum, ←finset.univ_product_univ, finset.sum_product, kronecker_map_apply] /-- `determinant` of `matrix.kronecker_map_bilinear`. diff --git a/src/data/matrix/notation.lean b/src/data/matrix/notation.lean index 881c6a1572fc0..dd35474e61c41 100644 --- a/src/data/matrix/notation.lean +++ b/src/data/matrix/notation.lean @@ -285,7 +285,7 @@ by { ext i, refine fin.cases _ _ i; simp [vec_mul_vec] } @[simp] lemma vec_mul_vec_cons (v : m' → α) (x : α) (w : fin n → α) : vec_mul_vec v (vec_cons x w) = λ i, v i • vec_cons x w := -by { ext i j, rw [vec_mul_vec, pi.smul_apply, smul_eq_mul] } +by { ext i j, rw [vec_mul_vec_apply, pi.smul_apply, smul_eq_mul] } end vec_mul_vec diff --git a/src/data/matrix/pequiv.lean b/src/data/matrix/pequiv.lean index 43cd965d4ce0f..e2fa164cd32df 100644 --- a/src/data/matrix/pequiv.lean +++ b/src/data/matrix/pequiv.lean @@ -44,8 +44,13 @@ open_locale matrix /-- `to_matrix` returns a matrix containing ones and zeros. `f.to_matrix i j` is `1` if `f i = some j` and `0` otherwise -/ -def to_matrix [decidable_eq n] [has_zero α] [has_one α] (f : m ≃. n) : matrix m n α -| i j := if j ∈ f i then 1 else 0 +def to_matrix [decidable_eq n] [has_zero α] [has_one α] (f : m ≃. n) : matrix m n α := +of $ λ i j, if j ∈ f i then (1 : α) else 0 + +-- TODO: set as an equation lemma for `to_matrix`, see mathlib4#3024 +@[simp] +lemma to_matrix_apply [decidable_eq n] [has_zero α] [has_one α] (f : m ≃. n) (i j) : + to_matrix f i j = if j ∈ f i then (1 : α) else 0 := rfl lemma mul_matrix_apply [fintype m] [decidable_eq m] [semiring α] (f : l ≃. m) (M : matrix m n α) (i j) : (f.to_matrix ⬝ M) i j = option.cases_on (f i) 0 (λ fi, M fi j) := @@ -59,11 +64,11 @@ end lemma to_matrix_symm [decidable_eq m] [decidable_eq n] [has_zero α] [has_one α] (f : m ≃. n) : (f.symm.to_matrix : matrix n m α) = f.to_matrixᵀ := -by ext; simp only [transpose, mem_iff_mem f, to_matrix]; congr +by ext; simp only [transpose, mem_iff_mem f, to_matrix_apply]; congr @[simp] lemma to_matrix_refl [decidable_eq n] [has_zero α] [has_one α] : ((pequiv.refl n).to_matrix : matrix n n α) = 1 := -by ext; simp [to_matrix, one_apply]; congr +by ext; simp [to_matrix_apply, one_apply]; congr lemma matrix_mul_apply [fintype m] [semiring α] [decidable_eq n] (M : matrix l m α) (f : m ≃. n) (i j) : (M ⬝ f.to_matrix) i j = option.cases_on (f.symm j) 0 (λ fj, M i fj) := @@ -104,7 +109,7 @@ begin classical, assume f g, refine not_imp_not.1 _, - simp only [matrix.ext_iff.symm, to_matrix, pequiv.ext_iff, + simp only [matrix.ext_iff.symm, to_matrix_apply, pequiv.ext_iff, not_forall, exists_imp_distrib], assume i hi, use i, diff --git a/src/linear_algebra/matrix/circulant.lean b/src/linear_algebra/matrix/circulant.lean index 9a0702d86b4b4..6b603bff25d5b 100644 --- a/src/linear_algebra/matrix/circulant.lean +++ b/src/linear_algebra/matrix/circulant.lean @@ -39,9 +39,12 @@ open_locale matrix big_operators /-- Given the condition `[has_sub n]` and a vector `v : n → α`, we define `circulant v` to be the circulant matrix generated by `v` of type `matrix n n α`. The `(i,j)`th entry is defined to be `v (i - j)`. -/ +def circulant [has_sub n] (v : n → α) : matrix n n α := +of $ λ i j, v (i - j) + +-- TODO: set as an equation lemma for `circulant`, see mathlib4#3024 @[simp] -def circulant [has_sub n] (v : n → α) : matrix n n α -| i j := v (i - j) +lemma circulant_apply [has_sub n] (v : n → α) (i j) : circulant v i j = v (i - j) := rfl lemma circulant_col_zero_eq [add_group n] (v : n → α) (i : n) : circulant v i 0 = v i := congr_arg v (sub_zero _) @@ -108,7 +111,7 @@ lemma circulant_mul [semiring α] [fintype n] [add_group n] (v w : n → α) : circulant v ⬝ circulant w = circulant (mul_vec (circulant v) w) := begin ext i j, - simp only [mul_apply, mul_vec, circulant, dot_product], + simp only [mul_apply, mul_vec, circulant_apply, dot_product], refine fintype.sum_equiv (equiv.sub_right j) _ _ _, intro x, simp only [equiv.sub_right_apply, sub_sub_sub_cancel_right], @@ -125,7 +128,7 @@ lemma circulant_mul_comm circulant v ⬝ circulant w = circulant w ⬝ circulant v := begin ext i j, - simp only [mul_apply, circulant, mul_comm], + simp only [mul_apply, circulant_apply, mul_comm], refine fintype.sum_equiv ((equiv.sub_left i).trans (equiv.add_right j)) _ _ _, intro x, congr' 2, diff --git a/src/linear_algebra/matrix/mv_polynomial.lean b/src/linear_algebra/matrix/mv_polynomial.lean index a850ba8f9809b..8c02733b22a3b 100644 --- a/src/linear_algebra/matrix/mv_polynomial.lean +++ b/src/linear_algebra/matrix/mv_polynomial.lean @@ -25,8 +25,13 @@ namespace matrix variables (m n R) /-- The matrix with variable `X (i,j)` at location `(i,j)`. -/ -@[simp] noncomputable def mv_polynomial_X [comm_semiring R] : matrix m n (mv_polynomial (m × n) R) -| i j := mv_polynomial.X (i, j) +noncomputable def mv_polynomial_X [comm_semiring R] : matrix m n (mv_polynomial (m × n) R) := +of $ λ i j, mv_polynomial.X (i, j) + +-- TODO: set as an equation lemma for `mv_polynomial_X`, see mathlib4#3024 +@[simp] +lemma mv_polynomial_X_apply [comm_semiring R] (i j) : + mv_polynomial_X m n R i j = mv_polynomial.X (i, j) := rfl variables {m n R S} diff --git a/src/linear_algebra/matrix/symmetric.lean b/src/linear_algebra/matrix/symmetric.lean index 6d37f21ced05b..5b1c8525d673d 100644 --- a/src/linear_algebra/matrix/symmetric.lean +++ b/src/linear_algebra/matrix/symmetric.lean @@ -121,8 +121,8 @@ end lemma is_symm_from_blocks_iff {A : matrix m m α} {B : matrix m n α} {C : matrix n m α} {D : matrix n n α} : (A.from_blocks B C D).is_symm ↔ A.is_symm ∧ Bᵀ = C ∧ Cᵀ = B ∧ D.is_symm := -⟨λ h, ⟨congr_arg to_blocks₁₁ h, congr_arg to_blocks₂₁ h, - congr_arg to_blocks₁₂ h, congr_arg to_blocks₂₂ h⟩, +⟨λ h, ⟨(congr_arg to_blocks₁₁ h : _), (congr_arg to_blocks₂₁ h : _), + (congr_arg to_blocks₁₂ h : _), (congr_arg to_blocks₂₂ h : _)⟩, λ ⟨hA, hBC, hCB, hD⟩, is_symm.from_blocks hA hBC hD⟩ end matrix diff --git a/src/linear_algebra/matrix/to_lin.lean b/src/linear_algebra/matrix/to_lin.lean index af53c60451b29..d9b4d10c6e987 100644 --- a/src/linear_algebra/matrix/to_lin.lean +++ b/src/linear_algebra/matrix/to_lin.lean @@ -627,7 +627,7 @@ linear_map.ext $ matrix.to_lin_fin_two_prod_apply _ _ _ _ begin ext, rw [linear_map.to_matrix_apply, distrib_mul_action.to_linear_map_apply, linear_equiv.map_smul, - basis.repr_self, finsupp.smul_single_one, finsupp.single_eq_pi_single, matrix.diagonal, + basis.repr_self, finsupp.smul_single_one, finsupp.single_eq_pi_single, matrix.diagonal_apply, pi.single_apply], end diff --git a/src/number_theory/cyclotomic/discriminant.lean b/src/number_theory/cyclotomic/discriminant.lean index 03682a366e33f..f92f560b96935 100644 --- a/src/number_theory/cyclotomic/discriminant.lean +++ b/src/number_theory/cyclotomic/discriminant.lean @@ -166,7 +166,7 @@ begin 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, + simp only [discr, trace_matrix_apply, 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 }, diff --git a/src/ring_theory/discriminant.lean b/src/ring_theory/discriminant.lean index 7c47309444ba8..21a559720dab7 100644 --- a/src/ring_theory/discriminant.lean +++ b/src/ring_theory/discriminant.lean @@ -89,7 +89,7 @@ begin { ext i, have : ∀ j, (trace A B) (b i * b j) * g j = (trace A B) (((g j) • (b j)) * b i), { intro j, simp [mul_comm], }, - simp only [mul_vec, dot_product, trace_matrix, pi.zero_apply, trace_form_apply, + simp only [mul_vec, dot_product, trace_matrix_apply, pi.zero_apply, trace_form_apply, λ j, this j, ← linear_map.map_sum, ← sum_mul, hg, zero_mul, linear_map.map_zero] }, by_contra h, rw discr_def at h, @@ -128,9 +128,9 @@ begin { have := span_eq_top_of_linear_independent_of_card_eq_finrank b.linear_independent (finrank_eq_card_basis b).symm, classical, - rw [discr_def, trace_matrix_def], + rw [discr_def, trace_matrix], simp_rw [← basis.mk_apply b.linear_independent this.ge], - rw [← trace_matrix_def, trace_matrix_of_basis, ← bilin_form.nondegenerate_iff_det_ne_zero], + rw [← trace_matrix, trace_matrix_of_basis, ← bilin_form.nondegenerate_iff_det_ne_zero], exact trace_form_nondegenerate _ _ }, end diff --git a/src/ring_theory/trace.lean b/src/ring_theory/trace.lean index 10188e27f7fe2..5f16361c7da38 100644 --- a/src/ring_theory/trace.lean +++ b/src/ring_theory/trace.lean @@ -404,11 +404,14 @@ open finset /-- Given an `A`-algebra `B` and `b`, an `κ`-indexed family of elements of `B`, we define `trace_matrix A b` as the matrix whose `(i j)`-th element is the trace of `b i * b j`. -/ -@[simp] noncomputable -def trace_matrix (b : κ → B) : matrix κ κ A -| i j := trace_form A B (b i) (b j) +noncomputable +def trace_matrix (b : κ → B) : matrix κ κ A := +of $ λ i j, trace_form A B (b i) (b j) -lemma trace_matrix_def (b : κ → B) : trace_matrix A b = of (λ i j, trace_form A B (b i) (b j)) := +-- TODO: set as an equation lemma for `trace_matrix`, see mathlib4#3024 +@[simp] +lemma trace_matrix_apply (b : κ → B) (i j) : + trace_matrix A b i j = trace_form A B (b i) (b j) := rfl lemma trace_matrix_reindex {κ' : Type*} (b : basis κ A B) (f : κ ≃ κ') : @@ -421,7 +424,7 @@ lemma trace_matrix_of_matrix_vec_mul [fintype κ] (b : κ → B) (P : matrix κ trace_matrix A ((P.map (algebra_map A B)).vec_mul b) = Pᵀ ⬝ (trace_matrix A b) ⬝ P := begin ext α β, - rw [trace_matrix, vec_mul, dot_product, vec_mul, dot_product, matrix.mul_apply, + rw [trace_matrix_apply, vec_mul, dot_product, vec_mul, dot_product, matrix.mul_apply, bilin_form.sum_left, fintype.sum_congr _ _ (λ (i : κ), @bilin_form.sum_right _ _ _ _ _ _ _ _ (b i * P.map (algebra_map A B) i α) (λ (y : κ), b y * P.map (algebra_map A B) y β)), sum_comm], congr, ext x, @@ -448,7 +451,7 @@ lemma trace_matrix_of_basis [fintype κ] [decidable_eq κ] (b : basis κ A B) : trace_matrix A b = bilin_form.to_matrix b (trace_form A B) := begin ext i j, - rw [trace_matrix, trace_form_apply, trace_form_to_matrix] + rw [trace_matrix_apply, trace_form_apply, trace_form_to_matrix] end lemma trace_matrix_of_basis_mul_vec (b : basis ι A B) (z : B) : @@ -456,7 +459,7 @@ lemma trace_matrix_of_basis_mul_vec (b : basis ι A B) (z : B) : begin ext i, rw [← col_apply ((trace_matrix A b).mul_vec (b.equiv_fun z)) i unit.star, col_mul_vec, - matrix.mul_apply, trace_matrix_def], + matrix.mul_apply, trace_matrix], simp only [col_apply, trace_form_apply], conv_lhs { congr, skip, funext, @@ -475,8 +478,12 @@ variable (A) /-- `embeddings_matrix A C b : matrix κ (B →ₐ[A] C) C` is the matrix whose `(i, σ)` coefficient is `σ (b i)`. It is mostly useful for fields when `fintype.card κ = finrank A B` and `C` is algebraically closed. -/ -@[simp] def embeddings_matrix (b : κ → B) : matrix κ (B →ₐ[A] C) C -| i σ := σ (b i) +def embeddings_matrix (b : κ → B) : matrix κ (B →ₐ[A] C) C := +of $ λ i (σ : B →ₐ[A] C), σ (b i) + +-- TODO: set as an equation lemma for `embeddings_matrix`, see mathlib4#3024 +@[simp] lemma embeddings_matrix_apply (b : κ → B) (i) (σ : B →ₐ[A] C) : + embeddings_matrix A C b i σ = σ (b i) := rfl /-- `embeddings_matrix_reindex A C b e : matrix κ κ C` is the matrix whose `(i, j)` coefficient is `σⱼ (b i)`, where `σⱼ : B →ₐ[A] C` is the embedding corresponding to `j : κ` given by a diff --git a/src/topology/instances/matrix.lean b/src/topology/instances/matrix.lean index 8c5478730bea7..dd801b2f40563 100644 --- a/src/topology/instances/matrix.lean +++ b/src/topology/instances/matrix.lean @@ -247,7 +247,7 @@ lemma continuous.matrix_block_diagonal' [has_zero R] [decidable_eq l] {A : X → Π i, matrix (m' i) (n' i) R} (hA : continuous A) : continuous (λ x, block_diagonal' (A x)) := continuous_matrix $ λ ⟨i₁, i₂⟩ ⟨j₁, j₂⟩, begin - dsimp only [block_diagonal'], + dsimp only [block_diagonal'_apply'], split_ifs, { subst h, exact ((continuous_apply i₁).comp hA).matrix_elem i₂ j₂ }, From 5923c4c7df633a157c9b752570a41ba5e8399289 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Fri, 31 Mar 2023 11:19:35 +0000 Subject: [PATCH 0742/1291] feat(order/rel_iso): well-founded instances for subtypes (#18699) --- src/order/rel_iso/basic.lean | 9 +++++++++ 1 file changed, 9 insertions(+) diff --git a/src/order/rel_iso/basic.lean b/src/order/rel_iso/basic.lean index 6652b264bfe09..ab68f286cbf01 100644 --- a/src/order/rel_iso/basic.lean +++ b/src/order/rel_iso/basic.lean @@ -300,9 +300,18 @@ end protected theorem well_founded : ∀ (f : r ↪r s) (h : well_founded s), well_founded r | f ⟨H⟩ := ⟨λ a, f.acc _ (H _)⟩ +protected theorem is_well_founded (f : r ↪r s) [is_well_founded β s] : is_well_founded α r := +⟨f.well_founded is_well_founded.wf⟩ + protected theorem is_well_order : ∀ (f : r ↪r s) [is_well_order β s], is_well_order α r | f H := by exactI {wf := f.well_founded H.wf, ..f.is_strict_total_order} +instance _root_.subtype.well_founded_lt [has_lt α] [well_founded_lt α] (p : α → Prop) : + well_founded_lt (subtype p) := (subtype.rel_embedding (<) p).is_well_founded + +instance _root_.subtype.well_founded_gt [has_lt α] [well_founded_gt α] (p : α → Prop) : + well_founded_gt (subtype p) := (subtype.rel_embedding (>) p).is_well_founded + /-- `quotient.mk` as a relation homomorphism between the relation and the lift of a relation. -/ @[simps] def _root_.quotient.mk_rel_hom [setoid α] {r : α → α → Prop} (H) : r →r quotient.lift₂ r H := From 29cb56a7b35f72758b05a30490e1f10bd62c35c1 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 31 Mar 2023 14:30:26 +0000 Subject: [PATCH 0743/1291] chore(*): Miscellaneous lemmas (#18677) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit * `algebra.support`: `support n = univ` if `n ≠ 0`, `mul_support n = univ` if `n ≠ 1` * `data.int.char_zero`: `↑n = 1 ↔ n = 1` * `data.real.ennreal`: `of_real a.to_real = a ↔ a ≠ ⊤`, `(of_real a).to_real = a ↔ 0 ≤ a` * `data.set.basic`: `s ∩ {a | p a} = {a ∈ s | p a}` * `logic.function.basic`: `on_fun f g a b = f (g a) (g b)` * `order.conditionally_complete_lattice.basic`: Lemmas unfolding the definition of `Sup`/`Inf` on `with_top`/`with_bot` --- src/algebra/support.lean | 30 +++++++++++++++++++ src/data/int/char_zero.lean | 6 ++++ src/data/real/ennreal.lean | 6 ++++ src/data/set/basic.lean | 20 +++++++++++++ src/logic/function/basic.lean | 3 ++ .../conditionally_complete_lattice/basic.lean | 14 +++++++++ 6 files changed, 79 insertions(+) diff --git a/src/algebra/support.lean b/src/algebra/support.lean index f0cdd96b62fca..2ecb97cc95cca 100644 --- a/src/algebra/support.lean +++ b/src/algebra/support.lean @@ -201,6 +201,36 @@ mul_support_binop_subset (/) one_div_one f g end division_monoid +section zero_one +variables (R) [has_zero R] [has_one R] [ne_zero (1 : R)] + +@[simp] lemma support_one : support (1 : α → R) = univ := support_const one_ne_zero +@[simp] lemma mul_support_zero : mul_support (0 : α → R) = univ := mul_support_const zero_ne_one + +end zero_one + +section add_monoid_with_one +variables [add_monoid_with_one R] [char_zero R] {n : ℕ} + +lemma support_nat_cast (hn : n ≠ 0) : support (n : α → R) = univ := +support_const $ nat.cast_ne_zero.2 hn + +lemma mul_support_nat_cast (hn : n ≠ 1) : mul_support (n : α → R) = univ := +mul_support_const $ nat.cast_ne_one.2 hn + +end add_monoid_with_one + +section add_group_with_one +variables [add_group_with_one R] [char_zero R] {n : ℤ} + +lemma support_int_cast (hn : n ≠ 0) : support (n : α → R) = univ := +support_const $ int.cast_ne_zero.2 hn + +lemma mul_support_int_cast (hn : n ≠ 1) : mul_support (n : α → R) = univ := +mul_support_const $ int.cast_ne_one.2 hn + +end add_group_with_one + lemma support_smul [has_zero R] [has_zero M] [smul_with_zero R M] [no_zero_smul_divisors R M] (f : α → R) (g : α → M) : support (f • g) = support f ∩ support g := diff --git a/src/data/int/char_zero.lean b/src/data/int/char_zero.lean index 91d70f00d7956..a233cbca4adb6 100644 --- a/src/data/int/char_zero.lean +++ b/src/data/int/char_zero.lean @@ -37,6 +37,12 @@ theorem cast_injective [add_group_with_one α] [char_zero α] : function.injecti theorem cast_ne_zero [add_group_with_one α] [char_zero α] {n : ℤ} : (n : α) ≠ 0 ↔ n ≠ 0 := not_congr cast_eq_zero +@[simp] lemma cast_eq_one [add_group_with_one α] [char_zero α] {n : ℤ} : (n : α) = 1 ↔ n = 1 := +by rw [←cast_one, cast_inj] + +lemma cast_ne_one [add_group_with_one α] [char_zero α] {n : ℤ} : (n : α) ≠ 1 ↔ n ≠ 1 := +cast_eq_one.not + @[simp, norm_cast] theorem cast_div_char_zero {k : Type*} [division_ring k] [char_zero k] {m n : ℤ} (n_dvd : n ∣ m) : ((m / n : ℤ) : k) = m / n := diff --git a/src/data/real/ennreal.lean b/src/data/real/ennreal.lean index 3976f5ddff8c1..5d778bbb1602e 100644 --- a/src/data/real/ennreal.lean +++ b/src/data/real/ennreal.lean @@ -210,6 +210,12 @@ by rw [ennreal.to_real, nnreal.coe_eq_one, ennreal.to_nnreal_eq_one_iff] @[simp] lemma of_real_lt_top {r : ℝ} : ennreal.of_real r < ∞ := lt_top_iff_ne_top.2 of_real_ne_top @[simp] lemma top_ne_of_real {r : ℝ} : ∞ ≠ ennreal.of_real r := by simp [ennreal.of_real] +@[simp] lemma of_real_to_real_eq_iff : ennreal.of_real a.to_real = a ↔ a ≠ ⊤ := +⟨λ h, by { rw ←h, exact of_real_ne_top }, of_real_to_real⟩ + +@[simp] lemma to_real_of_real_eq_iff {a : ℝ} : (ennreal.of_real a).to_real = a ↔ 0 ≤ a := +⟨λ h, by { rw ←h, exact to_real_nonneg }, to_real_of_real⟩ + @[simp] lemma zero_ne_top : 0 ≠ ∞ := coe_ne_top @[simp] lemma top_ne_zero : ∞ ≠ 0 := top_ne_coe diff --git a/src/data/set/basic.lean b/src/data/set/basic.lean index 279f8b6e1c8b3..e6224ffb184f7 100644 --- a/src/data/set/basic.lean +++ b/src/data/set/basic.lean @@ -615,6 +615,10 @@ inter_eq_self_of_subset_right $ subset_union_left _ _ theorem union_inter_cancel_right {s t : set α} : (s ∪ t) ∩ t = t := inter_eq_self_of_subset_right $ subset_union_right _ _ +lemma inter_set_of_eq_sep (s : set α) (p : α → Prop) : s ∩ {a | p a} = {a ∈ s | p a} := rfl +lemma set_of_inter_eq_sep (p : α → Prop) (s : set α) : {a | p a} ∩ s = {a ∈ s | p a} := +inter_comm _ _ + /-! ### Distributivity laws -/ theorem inter_distrib_left (s t u : set α) : s ∩ (t ∪ u) = (s ∩ t) ∪ (s ∩ u) := @@ -783,6 +787,8 @@ theorem insert_eq (x : α) (s : set α) : insert x s = ({x} : set α) ∪ s := r @[simp] theorem singleton_subset_iff {a : α} {s : set α} : {a} ⊆ s ↔ a ∈ s := forall_eq +lemma singleton_subset_singleton : ({a} : set α) ⊆ {b} ↔ a = b := by simp + theorem set_compr_eq_eq_singleton {a : α} : {b | b = a} = {a} := rfl @[simp] theorem singleton_union : {a} ∪ s = insert a s := rfl @@ -965,6 +971,12 @@ by rw [disjoint_singleton_left, mem_singleton_iff] lemma subset_diff : s ⊆ t \ u ↔ s ⊆ t ∧ disjoint s u := le_iff_subset.symm.trans le_sdiff +lemma inter_diff_distrib_left (s t u : set α) : s ∩ (t \ u) = (s ∩ t) \ (s ∩ u) := +inf_sdiff_distrib_left _ _ _ + +lemma inter_diff_distrib_right (s t u : set α) : s \ t ∩ u = (s ∩ u) \ (t ∩ u) := +inf_sdiff_distrib_right _ _ _ + /-! ### Lemmas about complement -/ lemma compl_def (s : set α) : sᶜ = {x | x ∉ s} := rfl @@ -1227,10 +1239,18 @@ sdiff_inf_self_right _ _ @[simp] theorem diff_singleton_eq_self {a : α} {s : set α} (h : a ∉ s) : s \ {a} = s := sdiff_eq_self_iff_disjoint.2 $ by simp [h] +@[simp] lemma diff_singleton_ssubset {s : set α} {a : α} : s \ {a} ⊂ s ↔ a ∈ s := +sdiff_le.lt_iff_ne.trans $ sdiff_eq_left.not.trans $ by simp + @[simp] theorem insert_diff_singleton {a : α} {s : set α} : insert a (s \ {a}) = insert a s := by simp [insert_eq, union_diff_self, -union_singleton, -singleton_union] +lemma insert_diff_singleton_comm (hab : a ≠ b) (s : set α) : + insert a (s \ {b}) = insert a s \ {b} := +by simp_rw [←union_singleton, union_diff_distrib, + diff_singleton_eq_self (mem_singleton_iff.not.2 hab.symm)] + @[simp] lemma diff_self {s : set α} : s \ s = ∅ := sdiff_self lemma diff_diff_right_self (s t : set α) : s \ (s \ t) = s ∩ t := sdiff_sdiff_right_self diff --git a/src/logic/function/basic.lean b/src/logic/function/basic.lean index f3691cdc4c00e..beb9d534798ad 100644 --- a/src/logic/function/basic.lean +++ b/src/logic/function/basic.lean @@ -46,6 +46,9 @@ lemma const_injective [nonempty α] : injective (const α : β → α → β) := lemma id_def : @id α = λ x, x := rfl +@[simp] lemma on_fun_apply (f : β → β → γ) (g : α → β) (a b : α) : on_fun f g a b = f (g a) (g b) := +rfl + lemma hfunext {α α': Sort u} {β : α → Sort v} {β' : α' → Sort v} {f : Πa, β a} {f' : Πa, β' a} (hα : α = α') (h : ∀a a', a == a' → f a == f' a') : f == f' := begin diff --git a/src/order/conditionally_complete_lattice/basic.lean b/src/order/conditionally_complete_lattice/basic.lean index 8eea89f3d5511..1f9cbdbbe8cda 100644 --- a/src/order/conditionally_complete_lattice/basic.lean +++ b/src/order/conditionally_complete_lattice/basic.lean @@ -58,6 +58,20 @@ noncomputable instance {α : Type*} [has_Sup α] : has_Sup (with_bot α) := noncomputable instance {α : Type*} [preorder α] [has_Inf α] : has_Inf (with_bot α) := ⟨(@with_top.has_Sup αᵒᵈ _ _).Sup⟩ +lemma with_top.Sup_eq [preorder α] [has_Sup α] {s : set (with_top α)} (hs : ⊤ ∉ s) + (hs' : bdd_above (coe ⁻¹' s : set α)) : Sup s = ↑(Sup (coe ⁻¹' s) : α) := +(if_neg hs).trans $ if_pos hs' + +lemma with_top.Inf_eq [has_Inf α] {s : set (with_top α)} (hs : ¬ s ⊆ {⊤}) : + Inf s = ↑(Inf (coe ⁻¹' s) : α) := if_neg hs + +lemma with_bot.Inf_eq [preorder α] [has_Inf α] {s : set (with_bot α)} (hs : ⊥ ∉ s) + (hs' : bdd_below (coe ⁻¹' s : set α)) : Inf s = ↑(Inf (coe ⁻¹' s) : α) := +(if_neg hs).trans $ if_pos hs' + +lemma with_bot.Sup_eq [has_Sup α] {s : set (with_bot α)} (hs : ¬ s ⊆ {⊥}) : + Sup s = ↑(Sup (coe ⁻¹' s) : α) := if_neg hs + @[simp] theorem with_top.cInf_empty {α : Type*} [has_Inf α] : Inf (∅ : set (with_top α)) = ⊤ := if_pos $ set.empty_subset _ From 47b12e7f2502f14001f891ca87fbae2b4acaed3f Mon Sep 17 00:00:00 2001 From: Moritz Doll Date: Fri, 31 Mar 2023 14:30:27 +0000 Subject: [PATCH 0744/1291] chore(analysis/locally_convex/strong_topology): generalize to semilinear maps (#18679) This is needed to show that the space of C-linear maps is locally convex. Also generalizes the ring from `real` to a general ordered semiring, since the proofs didn't need the `real`s. Co-authored-by: Eric Wieser --- .../locally_convex/strong_topology.lean | 32 +++++++++++-------- 1 file changed, 19 insertions(+), 13 deletions(-) diff --git a/src/analysis/locally_convex/strong_topology.lean b/src/analysis/locally_convex/strong_topology.lean index 79306e7a3b0ac..a5e444a5729da 100644 --- a/src/analysis/locally_convex/strong_topology.lean +++ b/src/analysis/locally_convex/strong_topology.lean @@ -27,25 +27,30 @@ locally convex, bounded convergence open_locale topology uniform_convergence -variables {E F : Type*} +variables {R 𝕜₁ 𝕜₂ E F : Type*} namespace continuous_linear_map +variables [add_comm_group E] [topological_space E] + [add_comm_group F] [topological_space F] [topological_add_group F] + section general -variables [add_comm_group E] [module ℝ E] [topological_space E] - [add_comm_group F] [module ℝ F] [topological_space F] [topological_add_group F] - [has_continuous_const_smul ℝ F] [locally_convex_space ℝ F] +variables (R) +variables [ordered_semiring R] +variables [normed_field 𝕜₁] [normed_field 𝕜₂] [module 𝕜₁ E] [module 𝕜₂ F] {σ : 𝕜₁ →+* 𝕜₂} +variables [module R F] [has_continuous_const_smul R F] [locally_convex_space R F] + [smul_comm_class 𝕜₂ R F] lemma strong_topology.locally_convex_space (𝔖 : set (set E)) (h𝔖₁ : 𝔖.nonempty) (h𝔖₂ : directed_on (⊆) 𝔖) : - @locally_convex_space ℝ (E →L[ℝ] F) _ _ _ (strong_topology (ring_hom.id ℝ) F 𝔖) := + @locally_convex_space R (E →SL[σ] F) _ _ _ (strong_topology σ F 𝔖) := begin - letI : topological_space (E →L[ℝ] F) := strong_topology (ring_hom.id ℝ) F 𝔖, - haveI : topological_add_group (E →L[ℝ] F) := strong_topology.topological_add_group _ _ _, + letI : topological_space (E →SL[σ] F) := strong_topology σ F 𝔖, + haveI : topological_add_group (E →SL[σ] F) := strong_topology.topological_add_group _ _ _, refine locally_convex_space.of_basis_zero _ _ _ _ (strong_topology.has_basis_nhds_zero_of_basis _ _ _ h𝔖₁ h𝔖₂ - (locally_convex_space.convex_basis_zero ℝ F)) _, + (locally_convex_space.convex_basis_zero R F)) _, rintros ⟨S, V⟩ ⟨hS, hVmem, hVconvex⟩ f hf g hg a b ha hb hab x hx, exact hVconvex (hf x hx) (hg x hx) ha hb hab, end @@ -54,12 +59,13 @@ end general section bounded_sets -variables [add_comm_group E] [module ℝ E] [topological_space E] - [add_comm_group F] [module ℝ F] [topological_space F] [topological_add_group F] - [has_continuous_const_smul ℝ F] [locally_convex_space ℝ F] +variables [ordered_semiring R] +variables [normed_field 𝕜₁] [normed_field 𝕜₂] [module 𝕜₁ E] [module 𝕜₂ F] {σ : 𝕜₁ →+* 𝕜₂} +variables [module R F] [has_continuous_const_smul R F] [locally_convex_space R F] + [smul_comm_class 𝕜₂ R F] -instance : locally_convex_space ℝ (E →L[ℝ] F) := -strong_topology.locally_convex_space _ ⟨∅, bornology.is_vonN_bounded_empty ℝ E⟩ +instance : locally_convex_space R (E →SL[σ] F) := +strong_topology.locally_convex_space R _ ⟨∅, bornology.is_vonN_bounded_empty 𝕜₁ E⟩ (directed_on_of_sup_mem $ λ _ _, bornology.is_vonN_bounded.union) end bounded_sets From 172bf2812857f5e56938cc148b7a539f52f84ca9 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Fri, 31 Mar 2023 17:55:55 +0000 Subject: [PATCH 0745/1291] feat(data/pnat/basic): pnat is a well-order (#18700) --- src/data/pnat/basic.lean | 2 ++ 1 file changed, 2 insertions(+) diff --git a/src/data/pnat/basic.lean b/src/data/pnat/basic.lean index 95208af3aadf8..fdbc7a7551636 100644 --- a/src/data/pnat/basic.lean +++ b/src/data/pnat/basic.lean @@ -26,6 +26,8 @@ attribute [derive [add_left_cancel_semigroup, add_right_cancel_semigroup, add_co namespace pnat +instance : is_well_order ℕ+ (<) := { } + @[simp] lemma one_add_nat_pred (n : ℕ+) : 1 + n.nat_pred = n := by rw [nat_pred, add_tsub_cancel_iff_le.mpr $ show 1 ≤ (n : ℕ), from n.2] From 02e095b22be8a3731542041cdc1266b7f588377c Mon Sep 17 00:00:00 2001 From: Alex J Best Date: Fri, 31 Mar 2023 20:41:04 +0000 Subject: [PATCH 0746/1291] feat(measure_theory/group/geometry_of_numbers): Blichfeldt and Minkowski's theorems (#2819) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit I had a go at Minkowski's convex body theorem and it works thanks to @urkud 's pidgeonhole for measurable spaces. Co-authored-by: Yaël Dillies Co-authored-by: Xavier-François Roblot <46200072+xroblot@users.noreply.github.com> Co-authored-by: RemyDegenne --- docs/100.yaml | 2 + docs/references.bib | 6 ++ .../group/geometry_of_numbers.lean | 83 +++++++++++++++++++ src/measure_theory/measure/haar_lebesgue.lean | 23 +++++ 4 files changed, 114 insertions(+) create mode 100644 src/measure_theory/group/geometry_of_numbers.lean diff --git a/docs/100.yaml b/docs/100.yaml index ecb9fd4cf2075..607e31750e8e4 100644 --- a/docs/100.yaml +++ b/docs/100.yaml @@ -154,6 +154,8 @@ note : "In `pell.eq_pell`, `d` is defined to be `a*a - 1` for an arbitrary `a > 1`." 40: title : Minkowski’s Fundamental Theorem + decl : measure_theory.exists_ne_zero_mem_lattice_of_measure_mul_two_pow_lt_measure + author : Alex J. Best, Yaël Dillies 41: title : Puiseux’s Theorem 42: diff --git a/docs/references.bib b/docs/references.bib index 8840f08ce1fda..cd493023456fa 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -577,6 +577,12 @@ @InProceedings{ CL21 series = {CPP 2021} } +@Book{ clark_gon, + author = {Pete L. Clark}, + title = {Geometry of Numbers with Applications to Number Theory}, + url = {http://alpha.math.uga.edu/~pete/geometryofnumbers.pdf} +} + @Book{ conway2001, author = {Conway, J. H.}, title = {On numbers and games}, diff --git a/src/measure_theory/group/geometry_of_numbers.lean b/src/measure_theory/group/geometry_of_numbers.lean new file mode 100644 index 0000000000000..0af202df3f364 --- /dev/null +++ b/src/measure_theory/group/geometry_of_numbers.lean @@ -0,0 +1,83 @@ +/- +Copyright (c) 2021 Alex J. Best. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Alex J. Best +-/ +import analysis.convex.measure +import measure_theory.group.fundamental_domain +import measure_theory.measure.haar_lebesgue + +/-! +# Geometry of numbers + +In this file we prove some of the fundamental theorems in the geometry of numbers, as studied by +Hermann Minkowski. + +## Main results + +* `exists_pair_mem_lattice_not_disjoint_vadd`: Blichfeldt's principle, existence of two distinct + points in a subgroup such that the translates of a set by these two points are not disjoint when + the covolume of the subgroup is larger than the volume of the +* `exists_ne_zero_mem_lattice_of_measure_mul_two_pow_lt_measure`: Minkowski's theorem, existence of + a non-zero lattice point inside a convex symmetric domain of large enough volume. + +## TODO + +* Calculate the volume of the fundamental domain of a finite index subgroup +* Voronoi diagrams +* See [Pete L. Clark, *Abstract Geometry of Numbers: Linear Forms* (arXiv)](https://arxiv.org/abs/1405.2119) + for some more ideas. + +## References + +* [Pete L. Clark, *Geometry of Numbers with Applications to Number Theory*][clark_gon] p.28 +-/ + +namespace measure_theory + +open ennreal finite_dimensional measure_theory measure_theory.measure set +open_locale pointwise + +variables {E L : Type*} [measurable_space E] {μ : measure E} {F s : set E} + +/-- **Blichfeldt's Theorem**. If the volume of the set `s` is larger than the covolume of the +countable subgroup `L` of `E`, then there exists two distincts points `x, y ∈ L` such that `(x + s)` +and `(y + s)` are not disjoint. -/ +lemma exists_pair_mem_lattice_not_disjoint_vadd [add_comm_group L] [countable L] + [add_action L E] [measurable_space L] [has_measurable_vadd L E] [vadd_invariant_measure L E μ] + (fund : is_add_fundamental_domain L F μ) (hS : null_measurable_set s μ) (h : μ F < μ s) : + ∃ x y : L, x ≠ y ∧ ¬ disjoint (x +ᵥ s) (y +ᵥ s) := +begin + contrapose! h, + exact ((fund.measure_eq_tsum _).trans (measure_Union₀ (pairwise.mono h $ λ i j hij, + (hij.mono inf_le_left inf_le_left).ae_disjoint) $ λ _, + (hS.vadd _).inter fund.null_measurable_set).symm).trans_le + (measure_mono $ Union_subset $ λ _, inter_subset_right _ _), +end + +/-- The **Minkowksi Convex Body Theorem**. If `s` is a convex symmetric domain of `E` whose volume +is large enough compared to the covolume of a lattice `L` of `E`, then it contains a non-zero +lattice point of `L`. -/ +lemma exists_ne_zero_mem_lattice_of_measure_mul_two_pow_lt_measure [normed_add_comm_group E] + [normed_space ℝ E] [borel_space E] [finite_dimensional ℝ E] [is_add_haar_measure μ] + {L : add_subgroup E} [countable L] (fund : is_add_fundamental_domain L F μ) + (h : μ F * 2 ^ finrank ℝ E < μ s) (h_symm : ∀ x ∈ s, -x ∈ s) (h_conv : convex ℝ s) : + ∃ x ≠ 0, ((x : L) : E) ∈ s := +begin + have h_vol : μ F < μ ((2⁻¹ : ℝ) • s), + { rwa [add_haar_smul_of_nonneg μ (by norm_num : 0 ≤ (2 : ℝ)⁻¹) s, ←mul_lt_mul_right + (pow_ne_zero (finrank ℝ E) (two_ne_zero' _)) (pow_ne_top two_ne_top), mul_right_comm, + of_real_pow (by norm_num : 0 ≤ (2 : ℝ)⁻¹), ←of_real_inv_of_pos zero_lt_two, of_real_bit0, + of_real_one, ←mul_pow, ennreal.inv_mul_cancel two_ne_zero two_ne_top, one_pow, one_mul] }, + obtain ⟨x, y, hxy, h⟩ := exists_pair_mem_lattice_not_disjoint_vadd fund + ((h_conv.smul _).null_measurable_set _) h_vol, + obtain ⟨_, ⟨v, hv, rfl⟩, w, hw, hvw⟩ := not_disjoint_iff.mp h, + refine ⟨x - y, sub_ne_zero.2 hxy, _⟩, + rw mem_inv_smul_set_iff₀ (two_ne_zero' ℝ) at hv hw, + simp_rw [add_subgroup.vadd_def, vadd_eq_add, add_comm _ w, ←sub_eq_sub_iff_add_eq_add, + ←add_subgroup.coe_sub] at hvw, + rw [←hvw, ←inv_smul_smul₀ (two_ne_zero' ℝ) (_ - _), smul_sub, sub_eq_add_neg, smul_add], + refine h_conv hw (h_symm _ hv) _ _ _; norm_num, +end + +end measure_theory diff --git a/src/measure_theory/measure/haar_lebesgue.lean b/src/measure_theory/measure/haar_lebesgue.lean index cfc90c4456988..ebfbc9b84360a 100644 --- a/src/measure_theory/measure/haar_lebesgue.lean +++ b/src/measure_theory/measure/haar_lebesgue.lean @@ -349,6 +349,29 @@ begin measure_singleton] } end +lemma add_haar_smul_of_nonneg {r : ℝ} (hr : 0 ≤ r) (s : set E) : + μ (r • s) = ennreal.of_real (r ^ finrank ℝ E) * μ s := +by rw [add_haar_smul, abs_pow, abs_of_nonneg hr] + +variables {μ} {s : set E} + +-- Note: We might want to rename this once we acquire the lemma corresponding to +-- `measurable_set.const_smul` +lemma null_measurable_set.const_smul (hs : null_measurable_set s μ) (r : ℝ) : + null_measurable_set (r • s) μ := +begin + obtain rfl | hs' := s.eq_empty_or_nonempty, + { simp }, + obtain rfl | hr := eq_or_ne r 0, + { simpa [zero_smul_set hs'] using null_measurable_set_singleton _ }, + obtain ⟨t, ht, hst⟩ := hs, + refine ⟨_, ht.const_smul_of_ne_zero hr, _⟩, + rw ←measure_symm_diff_eq_zero_iff at ⊢ hst, + rw [←smul_set_symm_diff₀ hr, add_haar_smul μ, hst, mul_zero], +end + +variables (μ) + @[simp] lemma add_haar_image_homothety (x : E) (r : ℝ) (s : set E) : μ (affine_map.homothety x r '' s) = ennreal.of_real (abs (r ^ (finrank ℝ E))) * μ s := calc μ (affine_map.homothety x r '' s) = μ ((λ y, y + x) '' (r • ((λ y, y + (-x)) '' s))) : From 210657c4ea4a4a7b234392f70a3a2a83346dfa90 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Fri, 31 Mar 2023 22:20:08 +0000 Subject: [PATCH 0747/1291] refactor(order/well_founded): ditch `well_founded_iff_has_min'` and `well_founded_iff_has_max'` (#15071) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The predicate `x ≤ y → y = x` is no more convenient than `¬ x < y`. For this reason, we ditch `well_founded.well_founded_iff_has_min'` and `well_founded.well_founded_iff_has_max'` in favor of `well_founded.well_founded_iff_has_min` (or in some cases, just `well_founded.has_min`. We also remove the misplaced lemma `well_founded.eq_iff_not_lt_of_le`, and we golf the theorems that used the removed theorems. The lemma `well_founded.well_founded_iff_has_min` has a misleading name when applied on `well_founded (>)`, and mildly screws over dot notation and rewriting by virtue of using `>`, but a future refactor will fix this. --- src/algebra/lie/engel.lean | 7 ++----- src/linear_algebra/free_module/pid.lean | 18 ++++++++--------- src/order/basic.lean | 3 +++ src/order/compactly_generated.lean | 22 ++++++++++----------- src/order/order_iso_nat.lean | 18 +++++------------ src/order/well_founded.lean | 20 ------------------- src/ring_theory/artinian.lean | 17 ++++++++-------- src/ring_theory/ideal/associated_prime.lean | 2 +- src/ring_theory/noetherian.lean | 18 ++++++++--------- 9 files changed, 45 insertions(+), 80 deletions(-) diff --git a/src/algebra/lie/engel.lean b/src/algebra/lie/engel.lean index fa9f351cae10e..23045db7fcb71 100644 --- a/src/algebra/lie/engel.lean +++ b/src/algebra/lie/engel.lean @@ -259,14 +259,11 @@ begin exact nontrivial_max_triv_of_is_nilpotent R K (L' ⧸ K.to_lie_submodule), }, haveI _i5 : is_noetherian R L' := is_noetherian_of_surjective L _ (linear_map.range_range_restrict (to_endomorphism R L M)), - obtain ⟨K, hK₁, hK₂⟩ := - well_founded.well_founded_iff_has_max'.mp (lie_subalgebra.well_founded_of_noetherian R L') s hs, + obtain ⟨K, hK₁, hK₂⟩ := (lie_subalgebra.well_founded_of_noetherian R L').has_min s hs, have hK₃ : K = ⊤, { by_contra contra, obtain ⟨K', hK'₁, hK'₂⟩ := this K hK₁ contra, - specialize hK₂ K' hK'₁ (le_of_lt hK'₂), - replace hK'₂ := (ne_of_lt hK'₂).symm, - contradiction, }, + exact hK₂ K' hK'₁ hK'₂, }, exact hK₃ ▸ hK₁, end diff --git a/src/linear_algebra/free_module/pid.lean b/src/linear_algebra/free_module/pid.lean index d975fcb650678..9f54c0e2859e8 100644 --- a/src/linear_algebra/free_module/pid.lean +++ b/src/linear_algebra/free_module/pid.lean @@ -59,7 +59,7 @@ variables {ι : Type*} (b : basis ι R M) open submodule.is_principal submodule lemma eq_bot_of_generator_maximal_map_eq_zero (b : basis ι R M) {N : submodule R M} - {ϕ : M →ₗ[R] R} (hϕ : ∀ (ψ : M →ₗ[R] R), N.map ϕ ≤ N.map ψ → N.map ψ = N.map ϕ) + {ϕ : M →ₗ[R] R} (hϕ : ∀ (ψ : M →ₗ[R] R), ¬ N.map ϕ < N.map ψ) [(N.map ϕ).is_principal] (hgen : generator (N.map ϕ) = (0 : R)) : N = ⊥ := begin rw submodule.eq_bot_iff, @@ -67,13 +67,13 @@ begin refine b.ext_elem (λ i, _), rw (eq_bot_iff_generator_eq_zero _).mpr hgen at hϕ, rw [linear_equiv.map_zero, finsupp.zero_apply], - exact (submodule.eq_bot_iff _).mp (hϕ ((finsupp.lapply i) ∘ₗ ↑b.repr) bot_le) _ ⟨x, hx, rfl⟩ + exact (submodule.eq_bot_iff _).mp (not_bot_lt_iff.1 $ hϕ ((finsupp.lapply i) ∘ₗ ↑b.repr)) _ + ⟨x, hx, rfl⟩ end lemma eq_bot_of_generator_maximal_submodule_image_eq_zero {N O : submodule R M} (b : basis ι R O) (hNO : N ≤ O) - {ϕ : O →ₗ[R] R} (hϕ : ∀ (ψ : O →ₗ[R] R), ϕ.submodule_image N ≤ ψ.submodule_image N → - ψ.submodule_image N = ϕ.submodule_image N) + {ϕ : O →ₗ[R] R} (hϕ : ∀ (ψ : O →ₗ[R] R), ¬ ϕ.submodule_image N < ψ.submodule_image N) [(ϕ.submodule_image N).is_principal] (hgen : generator (ϕ.submodule_image N) = 0) : N = ⊥ := begin @@ -82,7 +82,7 @@ begin refine congr_arg coe (show (⟨x, hNO hx⟩ : O) = 0, from b.ext_elem (λ i, _)), rw (eq_bot_iff_generator_eq_zero _).mpr hgen at hϕ, rw [linear_equiv.map_zero, finsupp.zero_apply], - refine (submodule.eq_bot_iff _).mp (hϕ ((finsupp.lapply i) ∘ₗ ↑b.repr) bot_le) _ _, + refine (submodule.eq_bot_iff _).mp (not_bot_lt_iff.1 $ hϕ ((finsupp.lapply i) ∘ₗ ↑b.repr)) _ _, exact (linear_map.mem_submodule_image_of_le hNO).mpr ⟨x, hx, rfl⟩ end @@ -115,8 +115,7 @@ variables {M : Type*} [add_comm_group M] [module R M] {b : ι → M} open submodule.is_principal lemma generator_maximal_submodule_image_dvd {N O : submodule R M} (hNO : N ≤ O) - {ϕ : O →ₗ[R] R} (hϕ : ∀ (ψ : O →ₗ[R] R), ϕ.submodule_image N ≤ ψ.submodule_image N → - ψ.submodule_image N = ϕ.submodule_image N) + {ϕ : O →ₗ[R] R} (hϕ : ∀ (ψ : O →ₗ[R] R), ¬ ϕ.submodule_image N < ψ.submodule_image N) [(ϕ.submodule_image N).is_principal] (y : M) (yN : y ∈ N) (ϕy_eq : ϕ ⟨y, hNO yN⟩ = generator (ϕ.submodule_image N)) (ψ : O →ₗ[R] R) : generator (ϕ.submodule_image N) ∣ ψ ⟨y, hNO yN⟩ := @@ -144,7 +143,7 @@ begin refine le_antisymm (this.trans (le_of_eq _)) (ideal.span_singleton_le_span_singleton.mpr d_dvd_left), rw span_singleton_generator, - refine hϕ ψ' (le_trans _ this), + apply (le_trans _ this).eq_of_not_gt (hϕ ψ'), rw [← span_singleton_generator (ϕ.submodule_image N)], exact ideal.span_singleton_le_span_singleton.mpr d_dvd_left, { exact subset_span (mem_insert _ _) } @@ -175,8 +174,7 @@ lemma submodule.basis_of_pid_aux [finite ι] {O : Type*} [add_comm_group O] [mod begin -- Let `ϕ` be a maximal projection of `M` onto `R`, in the sense that there is -- no `ψ` whose image of `N` is larger than `ϕ`'s image of `N`. - have : ∃ ϕ : M →ₗ[R] R, ∀ (ψ : M →ₗ[R] R), - ϕ.submodule_image N ≤ ψ.submodule_image N → ψ.submodule_image N = ϕ.submodule_image N, + have : ∃ ϕ : M →ₗ[R] R, ∀ (ψ : M →ₗ[R] R), ¬ ϕ.submodule_image N < ψ.submodule_image N, { obtain ⟨P, P_eq, P_max⟩ := set_has_maximal_iff_noetherian.mpr (infer_instance : is_noetherian R R) _ (show (set.range (λ ψ : M →ₗ[R] R, ψ.submodule_image N)).nonempty, diff --git a/src/order/basic.lean b/src/order/basic.lean index 708b8e3bdf4e2..f868ff486d619 100644 --- a/src/order/basic.lean +++ b/src/order/basic.lean @@ -227,6 +227,9 @@ le_iff_lt_or_eq.trans or.comm lemma lt_iff_le_and_ne [partial_order α] {a b : α} : a < b ↔ a ≤ b ∧ a ≠ b := ⟨λ h, ⟨le_of_lt h, ne_of_lt h⟩, λ ⟨h1, h2⟩, h1.lt_of_ne h2⟩ +lemma eq_iff_not_lt_of_le {α} [partial_order α] {x y : α} : x ≤ y → y = x ↔ ¬ x < y := +by rw [lt_iff_le_and_ne, not_and, not_not, eq_comm] + -- See Note [decidable namespace] protected lemma decidable.eq_iff_le_not_lt [partial_order α] [@decidable_rel α (≤)] {a b : α} : a = b ↔ a ≤ b ∧ ¬ a < b := diff --git a/src/order/compactly_generated.lean b/src/order/compactly_generated.lean index 6263ce4ff2347..47ee173833898 100644 --- a/src/order/compactly_generated.lean +++ b/src/order/compactly_generated.lean @@ -190,18 +190,16 @@ end lemma well_founded.is_Sup_finite_compact (h : well_founded ((>) : α → α → Prop)) : is_Sup_finite_compact α := -begin - intros s, - let p : set α := { x | ∃ (t : finset α), ↑t ⊆ s ∧ t.sup id = x }, - have hp : p.nonempty, { use [⊥, ∅], simp, }, - obtain ⟨m, ⟨t, ⟨ht₁, ht₂⟩⟩, hm⟩ := well_founded.well_founded_iff_has_max'.mp h p hp, - use t, simp only [ht₁, ht₂, true_and], apply le_antisymm, - { apply Sup_le, intros y hy, classical, - have hy' : (insert y t).sup id ∈ p, - { use insert y t, simp, rw set.insert_subset, exact ⟨hy, ht₁⟩, }, - have hm' : m ≤ (insert y t).sup id, { rw ← ht₂, exact finset.sup_mono (t.subset_insert y), }, - rw ← hm _ hy' hm', simp, }, - { rw [← ht₂, finset.sup_id_eq_Sup], exact Sup_le_Sup ht₁, }, +λ s, begin + obtain ⟨m, ⟨t, ⟨ht₁, rfl⟩⟩, hm⟩ := well_founded.well_founded_iff_has_min.mp h + {x | ∃ t : finset α, ↑t ⊆ s ∧ t.sup id = x} ⟨⊥, ∅, by simp⟩, + refine ⟨t, ht₁, (Sup_le (λ y hy, _)).antisymm _⟩, + { classical, + rw eq_of_le_of_not_lt (finset.sup_mono (t.subset_insert y)) + (hm _ ⟨insert y t, by simp [set.insert_subset, hy, ht₁]⟩), + simp }, + { rw finset.sup_id_eq_Sup, + exact Sup_le_Sup ht₁ }, end lemma is_Sup_finite_compact.is_sup_closed_compact (h : is_Sup_finite_compact α) : diff --git a/src/order/order_iso_nat.lean b/src/order/order_iso_nat.lean index 85145f7acda7c..b1221a36b6acf 100644 --- a/src/order/order_iso_nat.lean +++ b/src/order/order_iso_nat.lean @@ -226,17 +226,9 @@ a (monotonic_sequence_limit_index a) lemma well_founded.supr_eq_monotonic_sequence_limit [complete_lattice α] (h : well_founded ((>) : α → α → Prop)) (a : ℕ →o α) : supr a = monotonic_sequence_limit a := begin - suffices : (⨆ (m : ℕ), a m) ≤ monotonic_sequence_limit a, - { exact le_antisymm this (le_supr a _), }, - apply supr_le, - intros m, - by_cases hm : m ≤ monotonic_sequence_limit_index a, - { exact a.monotone hm, }, - { replace hm := le_of_not_le hm, - let S := { n | ∀ m, n ≤ m → a n = a m }, - have hInf : Inf S ∈ S, - { refine nat.Inf_mem _, rw well_founded.monotone_chain_condition at h, exact h a, }, - change Inf S ≤ m at hm, - change a m ≤ a (Inf S), - rw hInf m hm, }, + apply (supr_le (λ m, _)).antisymm (le_supr a _), + cases le_or_lt m (monotonic_sequence_limit_index a) with hm hm, + { exact a.monotone hm }, + { cases well_founded.monotone_chain_condition'.1 h a with n hn, + exact (nat.Inf_mem ⟨n, λ k hk, (a.mono hk).eq_of_not_lt (hn k hk)⟩ m hm.le).ge } end diff --git a/src/order/well_founded.lean b/src/order/well_founded.lean index 46f5f29ab47da..fb82de79a2a55 100644 --- a/src/order/well_founded.lean +++ b/src/order/well_founded.lean @@ -72,26 +72,6 @@ begin exact hm' y hy' hy end -lemma eq_iff_not_lt_of_le {α} [partial_order α] {x y : α} : x ≤ y → y = x ↔ ¬ x < y := -begin - split, - { intros xle nge, - cases le_not_le_of_lt nge, - rw xle left at nge, - exact lt_irrefl x nge }, - { intros ngt xle, - contrapose! ngt, - exact lt_of_le_of_ne xle (ne.symm ngt) } -end - -theorem well_founded_iff_has_max' [partial_order α] : (well_founded ((>) : α → α → Prop) ↔ - ∀ (p : set α), p.nonempty → ∃ m ∈ p, ∀ x ∈ p, m ≤ x → x = m) := -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' αᵒᵈ _ - open set /-- The supremum of a bounded, well-founded order -/ protected noncomputable def sup {r : α → α → Prop} (wf : well_founded r) (s : set α) diff --git a/src/ring_theory/artinian.lean b/src/ring_theory/artinian.lean index e720b4146490f..ed598587452c9 100644 --- a/src/ring_theory/artinian.lean +++ b/src/ring_theory/artinian.lean @@ -181,12 +181,11 @@ end /-- A module is Artinian iff every nonempty set of submodules has a minimal submodule among them. -/ theorem set_has_minimal_iff_artinian : - (∀ a : set $ submodule R M, a.nonempty → ∃ M' ∈ a, ∀ I ∈ a, I ≤ M' → I = M') ↔ - is_artinian R M := -by rw [is_artinian_iff_well_founded, well_founded.well_founded_iff_has_min'] + (∀ a : set $ submodule R M, a.nonempty → ∃ M' ∈ a, ∀ I ∈ a, ¬ I < M') ↔ is_artinian R M := +by rw [is_artinian_iff_well_founded, well_founded.well_founded_iff_has_min] theorem is_artinian.set_has_minimal [is_artinian R M] (a : set $ submodule R M) (ha : a.nonempty) : - ∃ M' ∈ a, ∀ I ∈ a, I ≤ M' → I = M' := + ∃ M' ∈ a, ∀ I ∈ a, ¬ I < M' := set_has_minimal_iff_artinian.mpr ‹_› a ha /-- A module is Artinian iff every decreasing chain of submodules stabilizes. -/ @@ -414,19 +413,19 @@ begin simpa only [this, top_smul, ideal.zero_eq_bot] using hJ }, by_contradiction hJ, change J ≠ ⊤ at hJ, rcases is_artinian.set_has_minimal {J' : ideal R | J < J'} ⟨⊤, hJ.lt_top⟩ - with ⟨J', hJJ' : J < J', hJ' : ∀ I, J < I → I ≤ J' → I = J'⟩, + with ⟨J', hJJ' : J < J', hJ' : ∀ I, J < I → ¬ I < J'⟩, rcases set_like.exists_of_lt hJJ' with ⟨x, hxJ', hxJ⟩, obtain rfl : J ⊔ ideal.span {x} = J', - { refine hJ' (J ⊔ ideal.span {x}) _ _, + { apply eq_of_le_of_not_lt _ (hJ' (J ⊔ ideal.span {x}) _), + { exact (sup_le hJJ'.le (span_le.2 (singleton_subset_iff.2 hxJ'))) }, { rw set_like.lt_iff_le_and_exists, - exact ⟨le_sup_left, ⟨x, mem_sup_right (mem_span_singleton_self x), hxJ⟩⟩ }, - { exact (sup_le hJJ'.le (span_le.2 (singleton_subset_iff.2 hxJ'))) } }, + exact ⟨le_sup_left, ⟨x, mem_sup_right (mem_span_singleton_self x), hxJ⟩⟩ } }, have : J ⊔ Jac • ideal.span {x} ≤ J ⊔ ideal.span {x}, from sup_le_sup_left (smul_le.2 (λ _ _ _, submodule.smul_mem _ _)) _, have : Jac * ideal.span {x} ≤ J, --Need version 4 of Nakayamas lemma on Stacks { classical, by_contradiction H, refine H (smul_sup_le_of_le_smul_of_le_jacobson_bot - (fg_span_singleton _) le_rfl (hJ' _ _ this).ge), + (fg_span_singleton _) le_rfl (this.eq_of_not_lt (hJ' _ _)).ge), exact lt_of_le_of_ne le_sup_left (λ h, H $ h.symm ▸ le_sup_right) }, have : ideal.span {x} * Jac ^ (n + 1) ≤ ⊥, calc ideal.span {x} * Jac ^ (n + 1) = ideal.span {x} * Jac * Jac ^ n : diff --git a/src/ring_theory/ideal/associated_prime.lean b/src/ring_theory/ideal/associated_prime.lean index 5ed07421fc294..c85ea9e010818 100644 --- a/src/ring_theory/ideal/associated_prime.lean +++ b/src/ring_theory/ideal/associated_prime.lean @@ -95,7 +95,7 @@ begin rw [smul_comm, hc, smul_zero] }, have H₂ : (submodule.span R {a • y}).annihilator ≠ ⊤, { rwa [ne.def, submodule.annihilator_eq_top_iff, submodule.span_singleton_eq_bot] }, - rwa [← h₃ (R ∙ a • y).annihilator ⟨l.trans H₁, H₂, _, rfl⟩ H₁, + rwa [H₁.eq_of_not_lt (h₃ (R ∙ a • y).annihilator ⟨l.trans H₁, H₂, _, rfl⟩), submodule.mem_annihilator_span_singleton, smul_comm, smul_smul] end diff --git a/src/ring_theory/noetherian.lean b/src/ring_theory/noetherian.lean index d2110c7afeeca..a0067dc2d6d6f 100644 --- a/src/ring_theory/noetherian.lean +++ b/src/ring_theory/noetherian.lean @@ -253,20 +253,20 @@ begin { intro H, constructor, intro N, - obtain ⟨⟨N₀, h₁⟩, e : N₀ ≤ N, h₂⟩ := well_founded.well_founded_iff_has_max'.mp + obtain ⟨⟨N₀, h₁⟩, e : N₀ ≤ N, h₂⟩ := well_founded.has_min H { N' : α | N'.1 ≤ N } ⟨⟨⊥, submodule.fg_bot⟩, bot_le⟩, convert h₁, refine (e.antisymm _).symm, by_contra h₃, obtain ⟨x, hx₁ : x ∈ N, hx₂ : x ∉ N₀⟩ := set.not_subset.mp h₃, apply hx₂, - have := h₂ ⟨(R ∙ x) ⊔ N₀, _⟩ _ _, + have := eq_of_le_of_not_lt _ (h₂ ⟨(R ∙ x) ⊔ N₀, _⟩ _), { injection this with eq, - rw ← eq, + rw eq, exact (le_sup_left : (R ∙ x) ≤ (R ∙ x) ⊔ N₀) (submodule.mem_span_singleton_self _) }, { exact submodule.fg.sup ⟨{x}, by rw [finset.coe_singleton]⟩ h₁ }, - { exact sup_le ((submodule.span_singleton_le_iff_mem _ _).mpr hx₁) e }, - { show N₀ ≤ (R ∙ x) ⊔ N₀, from le_sup_right } } + { show N₀ ≤ (R ∙ x) ⊔ N₀, from le_sup_right }, + { exact sup_le ((submodule.span_singleton_le_iff_mem _ _).mpr hx₁) e } } end variables (R M) @@ -280,14 +280,12 @@ variables {R M} /-- A module is Noetherian iff every nonempty set of submodules has a maximal submodule among them. -/ theorem set_has_maximal_iff_noetherian : - (∀ a : set $ submodule R M, a.nonempty → ∃ M' ∈ a, ∀ I ∈ a, M' ≤ I → I = M') ↔ - is_noetherian R M := -by rw [is_noetherian_iff_well_founded, well_founded.well_founded_iff_has_max'] + (∀ a : set $ submodule R M, a.nonempty → ∃ M' ∈ a, ∀ I ∈ a, ¬ M' < I) ↔ is_noetherian R M := +by rw [is_noetherian_iff_well_founded, well_founded.well_founded_iff_has_min] /-- A module is Noetherian iff every increasing chain of submodules stabilizes. -/ theorem monotone_stabilizes_iff_noetherian : - (∀ (f : ℕ →o submodule R M), ∃ n, ∀ m, n ≤ m → f n = f m) - ↔ is_noetherian R M := + (∀ (f : ℕ →o submodule R M), ∃ n, ∀ m, n ≤ m → f n = f m) ↔ is_noetherian R M := by rw [is_noetherian_iff_well_founded, well_founded.monotone_chain_condition] /-- If `∀ I > J, P I` implies `P J`, then `P` holds for all submodules. -/ From 5e526d18cea33550268dcbbddcb822d5cde40654 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 31 Mar 2023 22:20:10 +0000 Subject: [PATCH 0748/1291] =?UTF-8?q?feat(data/{set,finset}/pointwise):=20?= =?UTF-8?q?`a=20=E2=80=A2=20t=20=E2=8A=86=20s=20=E2=80=A2=20t`=20(#18697)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Eta expansion in the lemma statements is deliberate, to make the left and right lemmas more similar and allow further rewrites. Also additivise `finset.bUnion_smul_finset`, fix the name of `finset.smul_finset_mem_smul_finset` to `finset.smul_mem_smul_finset`, move `image2_swap`/`image₂_swap` further up the file to let them be used in earlier proofs. --- src/data/finset/n_ary.lean | 32 +++++++++++++--- src/data/finset/pointwise.lean | 64 +++++++++++++++++++++++++++++-- src/data/set/n_ary.lean | 45 ++++++++++++++-------- src/data/set/pointwise/basic.lean | 12 +++++- src/data/set/pointwise/smul.lean | 51 ++++++++++++++++++++++-- 5 files changed, 172 insertions(+), 32 deletions(-) diff --git a/src/data/finset/n_ary.lean b/src/data/finset/n_ary.lean index 80345cf7443ca..3a0db131eab7b 100644 --- a/src/data/finset/n_ary.lean +++ b/src/data/finset/n_ary.lean @@ -75,7 +75,7 @@ image₂_subset hs subset.rfl lemma image_subset_image₂_left (hb : b ∈ t) : s.image (λ a, f a b) ⊆ image₂ f s t := image_subset_iff.2 $ λ a ha, mem_image₂_of_mem ha hb -lemma image_subset_image₂_right (ha : a ∈ s) : t.image (f a) ⊆ image₂ f s t := +lemma image_subset_image₂_right (ha : a ∈ s) : t.image (λ b, f a b) ⊆ image₂ f s t := image_subset_iff.2 $ λ 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) := @@ -84,6 +84,12 @@ 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 +lemma image₂_subset_iff_left : image₂ f s t ⊆ u ↔ ∀ a ∈ s, t.image (λ b, f a b) ⊆ u := +by simp_rw [image₂_subset_iff, image_subset_iff] + +lemma image₂_subset_iff_right : image₂ f s t ⊆ u ↔ ∀ b ∈ t, s.image (λ a, f a b) ⊆ u := +by simp_rw [image₂_subset_iff, image_subset_iff, @forall₂_swap α] + @[simp] lemma image₂_nonempty_iff : (image₂ f s t).nonempty ↔ s.nonempty ∧ t.nonempty := by { rw [←coe_nonempty, coe_image₂], exact image2_nonempty_iff } @@ -113,6 +119,14 @@ 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 } +@[simp] lemma image₂_insert_left [decidable_eq α] : + image₂ f (insert a s) t = t.image (λ b, f a b) ∪ image₂ f s t := +coe_injective $ by { push_cast, exact image2_insert_left } + +@[simp] lemma image₂_insert_right [decidable_eq β] : + image₂ f s (insert b t) = s.image (λ a, f a b) ∪ image₂ f s t := +coe_injective $ by { push_cast, exact image2_insert_right } + lemma image₂_inter_left [decidable_eq α] (hf : injective2 f) : image₂ f (s ∩ s') t = image₂ f s t ∩ image₂ f s' t := coe_injective $ by { push_cast, exact image2_inter_left hf } @@ -214,10 +228,6 @@ 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₂_mk_eq_product [decidable_eq α] [decidable_eq β] (s : finset α) (t : finset β) : image₂ prod.mk s t = s ×ˢ t := by ext; simp [prod.ext_iff] @@ -229,6 +239,10 @@ by { classical, rw [←image₂_mk_eq_product, image_image₂, curry] } @[simp] lemma image_uncurry_product (f : α → β → γ) (s : finset α) (t : finset β) : (s ×ˢ t).image (uncurry f) = image₂ f s t := by rw [←image₂_curry, curry_uncurry] +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 } @@ -346,6 +360,14 @@ by rw [image₂_singleton_right, funext h, image_id'] variables [decidable_eq α] [decidable_eq β] +lemma image₂_inter_union_subset_union : + image₂ f (s ∩ s') (t ∪ t') ⊆ image₂ f s t ∪ image₂ f s' t' := +coe_subset.1 $ by { push_cast, exact set.image2_inter_union_subset_union } + +lemma image₂_union_inter_subset_union : + image₂ f (s ∪ s') (t ∩ t') ⊆ image₂ f s t ∪ image₂ f s' t' := +coe_subset.1 $ by { push_cast, exact set.image2_union_inter_subset_union } + lemma image₂_inter_union_subset {f : α → α → β} {s t : finset α} (hf : ∀ a b, f a b = f b a) : image₂ f (s ∩ t) (s ∪ t) ⊆ image₂ f s t := coe_subset.1 $ by { push_cast, exact image2_inter_union_subset hf } diff --git a/src/data/finset/pointwise.lean b/src/data/finset/pointwise.lean index d4833b70a33cb..03063a1ed2bc6 100644 --- a/src/data/finset/pointwise.lean +++ b/src/data/finset/pointwise.lean @@ -54,7 +54,7 @@ finset multiplication, finset addition, pointwise addition, pointwise multiplica pointwise subtraction -/ -open function +open function mul_opposite open_locale big_operators pointwise variables {F α β γ : Type*} @@ -124,7 +124,9 @@ localized "attribute [instance] finset.has_inv finset.has_neg" in pointwise @[simp, to_additive] lemma inv_empty : (∅ : finset α)⁻¹ = ∅ := image_empty _ @[simp, to_additive] lemma inv_nonempty_iff : s⁻¹.nonempty ↔ s.nonempty := nonempty.image_iff _ -alias inv_nonempty_iff ↔ nonempty.inv nonempty.of_inv +alias inv_nonempty_iff ↔ nonempty.of_inv nonempty.inv + +attribute [to_additive] nonempty.inv nonempty.of_inv @[to_additive, mono] lemma inv_subset_inv (h : s ⊆ t) : s⁻¹ ⊆ t⁻¹ := image_subset_image h @@ -212,6 +214,10 @@ attribute [mono] add_subset_add image₂_inter_subset_left @[to_additive] lemma mul_inter_subset : s * (t₁ ∩ t₂) ⊆ s * t₁ ∩ (s * t₂) := image₂_inter_subset_right +@[to_additive] lemma inter_mul_union_subset_union : s₁ ∩ s₂ * (t₁ ∪ t₂) ⊆ (s₁ * t₁) ∪ (s₂ * t₂) := +image₂_inter_union_subset_union +@[to_additive] lemma union_mul_inter_subset_union : (s₁ ∪ s₂) * (t₁ ∩ t₂) ⊆ (s₁ * t₁) ∪ (s₂ * t₂) := +image₂_union_inter_subset_union /-- 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'`. -/ @@ -294,6 +300,10 @@ attribute [mono] sub_subset_sub image₂_inter_subset_left @[to_additive] lemma div_inter_subset : s / (t₁ ∩ t₂) ⊆ s / t₁ ∩ (s / t₂) := image₂_inter_subset_right +@[to_additive] lemma inter_div_union_subset_union : (s₁ ∩ s₂) / (t₁ ∪ t₂) ⊆ (s₁ / t₁) ∪ (s₂ / t₂) := +image₂_inter_union_subset_union +@[to_additive] lemma union_div_inter_subset_union : (s₁ ∪ s₂) / (t₁ ∩ t₂) ⊆ (s₁ / t₁) ∪ (s₂ / t₂) := +image₂_union_inter_subset_union /-- 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'`. -/ @@ -711,6 +721,12 @@ image₂_union_left image₂_inter_subset_left @[to_additive] lemma smul_inter_subset : s • (t₁ ∩ t₂) ⊆ s • t₁ ∩ s • t₂ := image₂_inter_subset_right +@[to_additive] lemma inter_smul_union_subset_union [decidable_eq α] : + (s₁ ∩ s₂) • (t₁ ∪ t₂) ⊆ (s₁ • t₁) ∪ (s₂ • t₂) := +image₂_inter_union_subset_union +@[to_additive] lemma union_smul_inter_subset_union [decidable_eq α] : + (s₁ ∪ s₂) • (t₁ ∩ t₂) ⊆ (s₁ • t₁) ∪ (s₂ • t₂) := +image₂_union_inter_subset_union /-- 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'`. -/ @@ -806,7 +822,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 (a : α) (s : finset β) : (↑(a • s) : set β) = a • s := coe_image -@[to_additive] lemma smul_finset_mem_smul_finset : b ∈ s → a • b ∈ a • s := mem_image_of_mem _ +@[to_additive] lemma smul_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 _ @@ -829,7 +845,11 @@ lemma smul_finset_singleton (b : β) : a • ({b} : finset β) = {a • b} := im @[to_additive] lemma smul_finset_inter_subset : a • (s₁ ∩ s₂) ⊆ a • s₁ ∩ (a • s₂) := image_inter_subset _ _ _ -@[simp] lemma bUnion_smul_finset (s : finset α) (t : finset β) : s.bUnion (• t) = s • t := +@[to_additive] lemma smul_finset_subset_smul {s : finset α} : a ∈ s → a • t ⊆ s • t := +image_subset_image₂_right + +@[simp, to_additive] lemma bUnion_smul_finset (s : finset α) (t : finset β) : + s.bUnion (• t) = s • t := bUnion_image_left end has_smul @@ -936,6 +956,42 @@ coe_injective.no_zero_smul_divisors _ coe_zero coe_smul_finset end instances +section has_smul +variables [decidable_eq β] [decidable_eq γ] [has_smul αᵐᵒᵖ β] [has_smul β γ] [has_smul α γ] + +-- TODO: replace hypothesis and conclusion with a typeclass +@[to_additive] lemma op_smul_finset_smul_eq_smul_smul_finset (a : α) (s : finset β) (t : finset γ) + (h : ∀ (a : α) (b : β) (c : γ), (op a • b) • c = b • a • c) : + (op a • s) • t = s • a • t := +by { ext, simp [mem_smul, mem_smul_finset, h] } + +end has_smul + +section has_mul +variables [has_mul α] [decidable_eq α] {s t u : finset α} {a : α} + +@[to_additive] lemma op_smul_finset_subset_mul : a ∈ t → op a • s ⊆ s * t := +image_subset_image₂_left + +@[simp, to_additive] lemma bUnion_op_smul_finset (s t : finset α) : + t.bUnion (λ a, op a • s) = s * t := +bUnion_image_right + +@[to_additive] lemma mul_subset_iff_left : s * t ⊆ u ↔ ∀ a ∈ s, a • t ⊆ u := image₂_subset_iff_left +@[to_additive] lemma mul_subset_iff_right : s * t ⊆ u ↔ ∀ b ∈ t, op b • s ⊆ u := +image₂_subset_iff_right + +end has_mul + +section semigroup +variables [semigroup α] [decidable_eq α] + +@[to_additive] lemma op_smul_finset_mul_eq_mul_smul_finset (a : α) (s : finset α) (t : finset α) : + (op a • s) * t = s * a • t := +op_smul_finset_smul_eq_smul_smul_finset _ _ _ $ λ _ _ _, mul_assoc _ _ _ + +end semigroup + section left_cancel_semigroup variables [left_cancel_semigroup α] [decidable_eq α] (s t : finset α) (a : α) diff --git a/src/data/set/n_ary.lean b/src/data/set/n_ary.lean index 27c35f92b03af..0ebf7b6ec2f87 100644 --- a/src/data/set/n_ary.lean +++ b/src/data/set/n_ary.lean @@ -67,6 +67,12 @@ lemma forall_image2_iff {p : γ → Prop} : image2 f s t ⊆ u ↔ ∀ (x ∈ s) (y ∈ t), f x y ∈ u := forall_image2_iff +lemma image2_subset_iff_left : image2 f s t ⊆ u ↔ ∀ a ∈ s, (λ b, f a b) '' t ⊆ u := +by simp_rw [image2_subset_iff, image_subset_iff, subset_def, mem_preimage] + +lemma image2_subset_iff_right : image2 f s t ⊆ u ↔ ∀ b ∈ t, (λ a, f a b) '' s ⊆ u := +by simp_rw [image2_subset_iff, image_subset_iff, subset_def, mem_preimage, @forall₂_swap α] + variables (f) @[simp] lemma image_prod : (λ x : α × β, f x.1 x.2) '' s ×ˢ t = image2 f s t := @@ -83,6 +89,9 @@ image_prod _ image2 (λ a b, f (a, b)) s t = f '' s ×ˢ t := by simp [←image_uncurry_prod, uncurry] +lemma image2_swap (s : set α) (t : set β) : image2 f s t = image2 (λ a b, f b a) t s := +by { ext, split; rintro ⟨a, b, ha, hb, rfl⟩; refine ⟨b, a, hb, ha, rfl⟩ } + variables {f} lemma image2_union_left : image2 f (s ∪ s') t = image2 f s t ∪ image2 f s' t := @@ -95,13 +104,7 @@ begin end lemma image2_union_right : image2 f s (t ∪ t') = image2 f s t ∪ image2 f s t' := -begin - ext c, - split, - { rintro ⟨a, b, ha, h1b|h2b, rfl⟩;[left, right]; exact ⟨_, _, ‹_›, ‹_›, rfl⟩ }, - { rintro (⟨_, _, _, _, rfl⟩|⟨_, _, _, _, rfl⟩); refine ⟨_, _, ‹_›, _, rfl⟩; - simp [mem_union, *] } -end +by rw [←image2_swap, image2_union_left, image2_swap f, image2_swap f] lemma image2_inter_left (hf : injective2 f) : image2 f (s ∩ s') t = image2 f s t ∩ image2 f s' t := by simp_rw [←image_uncurry_prod, inter_prod, image_inter hf.uncurry] @@ -138,6 +141,12 @@ by { rintro _ ⟨a, b, ha, ⟨h1b, h2b⟩, rfl⟩, split; exact ⟨_, _, ‹_› lemma image2_singleton : image2 f {a} {b} = {f a b} := by simp +@[simp] lemma image2_insert_left : image2 f (insert a s) t = (λ b, f a b) '' t ∪ image2 f s t := +by rw [insert_eq, image2_union_left, image2_singleton_left] + +@[simp] lemma image2_insert_right : image2 f s (insert b t) = (λ a, f a b) '' s ∪ image2 f s t := +by rw [insert_eq, image2_union_right, image2_singleton_right] + @[congr] lemma image2_congr (h : ∀ (a ∈ s) (b ∈ t), f a b = f' a b) : image2 f s t = image2 f' s t := by { ext, split; rintro ⟨a, b, ha, hb, rfl⟩; refine ⟨a, b, ha, hb, by rw h a ha b hb⟩ } @@ -208,10 +217,6 @@ begin { rintro ⟨a, b, ha, hb, rfl⟩, refine ⟨a, _, ha, ⟨b, hb, rfl⟩, rfl⟩ } end -lemma image2_swap (f : α → β → γ) (s : set α) (t : set β) : - image2 f s t = image2 (λ a b, f b a) t s := -by { ext, split; rintro ⟨a, b, ha, hb, rfl⟩; refine ⟨b, a, hb, ha, rfl⟩ } - @[simp] lemma image2_left (h : t.nonempty) : image2 (λ x y, x) s t = s := by simp [nonempty_def.mp h, ext_iff] @@ -339,14 +344,20 @@ lemma image2_right_identity {f : α → β → α} {b : β} (h : ∀ a, f a b = image2 f s {b} = s := by rw [image2_singleton_right, funext h, image_id'] +lemma image2_inter_union_subset_union : + image2 f (s ∩ s') (t ∪ t') ⊆ image2 f s t ∪ image2 f s' t' := +by { rw image2_union_right, exact union_subset_union + (image2_subset_right $ inter_subset_left _ _) (image2_subset_right $ inter_subset_right _ _) } + +lemma image2_union_inter_subset_union : + image2 f (s ∪ s') (t ∩ t') ⊆ image2 f s t ∪ image2 f s' t' := +by { rw image2_union_left, exact union_subset_union + (image2_subset_left $ inter_subset_left _ _) (image2_subset_left $ inter_subset_right _ _) } + lemma image2_inter_union_subset {f : α → α → β} {s t : set α} (hf : ∀ a b, f a b = f b a) : image2 f (s ∩ t) (s ∪ t) ⊆ image2 f s t := -begin - rintro _ ⟨a, b, ha, hb | hb, rfl⟩, - { rw hf, - exact mem_image2_of_mem hb ha.2 }, - { exact mem_image2_of_mem ha.1 hb } -end +by { rw inter_comm, + exact image2_inter_union_subset_union.trans (union_subset (image2_comm hf).subset subset.rfl) } lemma image2_union_inter_subset {f : α → α → β} {s t : set α} (hf : ∀ a b, f a b = f b a) : image2 f (s ∪ t) (s ∩ t) ⊆ image2 f s t := diff --git a/src/data/set/pointwise/basic.lean b/src/data/set/pointwise/basic.lean index aa0c9c0ac8a2f..5b100178a201a 100644 --- a/src/data/set/pointwise/basic.lean +++ b/src/data/set/pointwise/basic.lean @@ -223,6 +223,10 @@ attribute [mono] add_subset_add 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_union_subset_union : s₁ ∩ s₂ * (t₁ ∪ t₂) ⊆ (s₁ * t₁) ∪ (s₂ * t₂) := +image2_inter_union_subset_union +@[to_additive] lemma union_mul_inter_subset_union : (s₁ ∪ s₂) * (t₁ ∩ t₂) ⊆ (s₁ * t₁) ∪ (s₂ * t₂) := +image2_union_inter_subset_union @[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 _ @@ -324,6 +328,10 @@ attribute [mono] sub_subset_sub 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_union_subset_union : s₁ ∩ s₂ / (t₁ ∪ t₂) ⊆ (s₁ / t₁) ∪ (s₂ / t₂) := +image2_inter_union_subset_union +@[to_additive] lemma union_div_inter_subset_union : (s₁ ∪ s₂) / (t₁ ∩ t₂) ⊆ (s₁ / t₁) ∪ (s₂ / t₂) := +image2_union_inter_subset_union @[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 _ @@ -412,8 +420,8 @@ variables [mul_one_class α] /-- `set α` is a `mul_one_class` under pointwise operations if `α` is. -/ @[to_additive "`set α` is an `add_zero_class` under pointwise operations if `α` is."] protected def mul_one_class : mul_one_class (set α) := -{ mul_one := λ s, by { simp only [← singleton_one, mul_singleton, mul_one, image_id'] }, - one_mul := λ s, by { simp only [← singleton_one, singleton_mul, one_mul, image_id'] }, +{ mul_one := image2_right_identity mul_one, + one_mul := image2_left_identity one_mul, ..set.has_one, ..set.has_mul } localized "attribute [instance] set.mul_one_class set.add_zero_class set.semigroup set.add_semigroup diff --git a/src/data/set/pointwise/smul.lean b/src/data/set/pointwise/smul.lean index 75e8c8d504d26..2b807aa3215b2 100644 --- a/src/data/set/pointwise/smul.lean +++ b/src/data/set/pointwise/smul.lean @@ -38,7 +38,7 @@ Appropriate definitions and results are also transported to the additive theory -/ -open function +open function mul_opposite variables {F α β γ : Type*} @@ -108,6 +108,12 @@ attribute [mono] vadd_subset_vadd @[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_union_subset_union : + (s₁ ∩ s₂) • (t₁ ∪ t₂) ⊆ (s₁ • t₁) ∪ (s₂ • t₂) := +image2_inter_union_subset_union +@[to_additive] lemma union_smul_inter_subset_union : + (s₁ ∪ s₂) • (t₁ ∩ t₂) ⊆ (s₁ • t₁) ∪ (s₂ • t₂) := +image2_union_inter_subset_union @[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 _ @@ -143,6 +149,9 @@ 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 smul_set_subset_smul {s : set α} : a ∈ s → a • t ⊆ s • t := +image_subset_image2_right + @[simp, to_additive] lemma bUnion_smul_set (s : set α) (t : set β) : (⋃ a ∈ s, a • t) = s • t := Union_image_left _ @@ -192,12 +201,22 @@ image_Inter₂_subset _ _ end has_smul_set -variables {s s₁ s₂ : set α} {t t₁ t₂ : set β} {a : α} {b : β} +section has_mul +variables [has_mul α] {s t u : set α} {a : α} -@[simp, to_additive] lemma bUnion_op_smul_set [has_mul α] (s t : set α) : - (⋃ a ∈ t, mul_opposite.op a • s) = s * t := +@[to_additive] lemma op_smul_set_subset_mul : a ∈ t → op a • s ⊆ s * t := image_subset_image2_left + +@[simp, to_additive] lemma bUnion_op_smul_set (s t : set α) : (⋃ a ∈ t, op a • s) = s * t := Union_image_right _ +@[to_additive] lemma mul_subset_iff_left : s * t ⊆ u ↔ ∀ a ∈ s, a • t ⊆ u := image2_subset_iff_left +@[to_additive] lemma mul_subset_iff_right : s * t ⊆ u ↔ ∀ b ∈ t, op b • s ⊆ u := +image2_subset_iff_right + +end has_mul + +variables {s s₁ s₂ : set α} {t t₁ t₂ : set β} {a : α} {b : β} + @[to_additive] theorem range_smul_range {ι κ : Type*} [has_smul α β] (b : ι → α) (c : κ → β) : range b • range c = range (λ p : ι × κ, b p.1 • c p.2) := @@ -345,6 +364,10 @@ lemma union_vsub : (s₁ ∪ s₂) -ᵥ t = s₁ -ᵥ t ∪ (s₂ -ᵥ t) := ima 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 +lemma inter_vsub_union_subset_union : (s₁ ∩ s₂) -ᵥ (t₁ ∪ t₂) ⊆ (s₁ -ᵥ t₁) ∪ (s₂ -ᵥ t₂) := +image2_inter_union_subset_union +lemma union_vsub_inter_subset_union : (s₁ ∪ s₂) -ᵥ (t₁ ∩ t₂) ⊆ (s₁ -ᵥ t₁) ∪ (s₂ -ᵥ t₂) := +image2_union_inter_subset_union lemma Union_vsub_left_image : (⋃ a ∈ s, ((-ᵥ) a) '' t) = s -ᵥ t := Union_image_left _ lemma Union_vsub_right_image : (⋃ a ∈ t, (-ᵥ a) '' s) = s -ᵥ t := Union_image_right _ @@ -388,6 +411,17 @@ image_comm f '' (a • s) = f a • f '' s := image_comm $ map_mul _ _ +section has_smul +variables[has_smul αᵐᵒᵖ β] [has_smul β γ] [has_smul α γ] + +-- TODO: replace hypothesis and conclusion with a typeclass +@[to_additive] lemma op_smul_set_smul_eq_smul_smul_set (a : α) (s : set β) (t : set γ) + (h : ∀ (a : α) (b : β) (c : γ), (op a • b) • c = b • a • c) : + (op a • s) • t = s • a • t := +by { ext, simp [mem_smul, mem_smul_set, h] } + +end has_smul + section smul_with_zero variables [has_zero α] [has_zero β] [smul_with_zero α β] {s : set α} {t : set β} @@ -441,6 +475,15 @@ end end smul_with_zero +section semigroup +variables [semigroup α] + +@[to_additive] lemma op_smul_set_mul_eq_mul_smul_set (a : α) (s : set α) (t : set α) : + (op a • s) * t = s * a • t := +op_smul_set_smul_eq_smul_smul_set _ _ _ $ λ _ _ _, mul_assoc _ _ _ + +end semigroup + section left_cancel_semigroup variables [left_cancel_semigroup α] {s t : set α} From 474656fdf40ae1741dfffcdd7c685a0f198da61a Mon Sep 17 00:00:00 2001 From: Christopher Hoskin Date: Sat, 1 Apr 2023 01:35:41 +0000 Subject: [PATCH 0749/1291] chore(algebra/order/lattice_group): Golf (#18046) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Golf proofs and remove a duplicate lemma. Co-authored-by: Christopher Hoskin Co-authored-by: Yaël Dillies --- src/algebra/order/lattice_group.lean | 127 ++++++++------------------- 1 file changed, 38 insertions(+), 89 deletions(-) diff --git a/src/algebra/order/lattice_group.lean b/src/algebra/order/lattice_group.lean index 3993c527055ea..b1cd7376dd9a7 100644 --- a/src/algebra/order/lattice_group.lean +++ b/src/algebra/order/lattice_group.lean @@ -74,46 +74,37 @@ variables {α : Type u} [lattice α] [comm_group α] -- c + (a ⊔ b) = (c + a) ⊔ (c + b) @[to_additive] lemma mul_sup [covariant_class α α (*) (≤)] (a b c : α) : c * (a ⊔ b) = (c * a) ⊔ (c * b) := -begin - refine le_antisymm _ (by simp), - rw [← mul_le_mul_iff_left (c⁻¹), ← mul_assoc, inv_mul_self, one_mul], - exact sup_le (by simp) (by simp), -end +(order_iso.mul_left _).map_sup _ _ + +@[to_additive] +lemma sup_mul [covariant_class α α (*) (≤)] (a b c : α) : (a ⊔ b) * c = (a * c) ⊔ (b * c) := +(order_iso.mul_right _).map_sup _ _ @[to_additive] lemma mul_inf [covariant_class α α (*) (≤)] (a b c : α) : c * (a ⊓ b) = (c * a) ⊓ (c * b) := -begin - refine le_antisymm (by simp) _, - rw [← mul_le_mul_iff_left (c⁻¹), ← mul_assoc, inv_mul_self, one_mul], - exact le_inf (by simp) (by simp), -end +(order_iso.mul_left _).map_inf _ _ + +@[to_additive] +lemma inf_mul [covariant_class α α (*) (≤)] (a b c : α) : (a ⊓ b) * c = (a * c) ⊓ (b * c) := +(order_iso.mul_right _).map_inf _ _ -- Special case of Bourbaki A.VI.9 (2) -- -(a ⊔ b)=(-a) ⊓ (-b) @[to_additive] lemma inv_sup_eq_inv_inf_inv [covariant_class α α (*) (≤)] (a b : α) : (a ⊔ b)⁻¹ = a⁻¹ ⊓ b⁻¹ := -begin - apply le_antisymm, - { refine le_inf _ _, - { rw inv_le_inv_iff, exact le_sup_left, }, - { rw inv_le_inv_iff, exact le_sup_right, } }, - { rw [← inv_le_inv_iff, inv_inv], - refine sup_le _ _, - { rw ← inv_le_inv_iff, simp, }, - { rw ← inv_le_inv_iff, simp, } } -end +(order_iso.inv α).map_sup _ _ -- -(a ⊓ b) = -a ⊔ -b @[to_additive] lemma inv_inf_eq_sup_inv [covariant_class α α (*) (≤)] (a b : α) : (a ⊓ b)⁻¹ = a⁻¹ ⊔ b⁻¹ := -by rw [← inv_inv (a⁻¹ ⊔ b⁻¹), inv_sup_eq_inv_inf_inv a⁻¹ b⁻¹, inv_inv, inv_inv] +(order_iso.inv α).map_inf _ _ -- Bourbaki A.VI.10 Prop 7 -- a ⊓ b + (a ⊔ b) = a + b @[to_additive] lemma inf_mul_sup [covariant_class α α (*) (≤)] (a b : α) : (a ⊓ b) * (a ⊔ b) = a * b := calc (a ⊓ b) * (a ⊔ b) = (a ⊓ b) * ((a * b) * (b⁻¹ ⊔ a⁻¹)) : - by { rw mul_sup b⁻¹ a⁻¹ (a * b), simp, } + by rw [mul_sup b⁻¹ a⁻¹ (a * b), mul_inv_cancel_right, mul_inv_cancel_comm] ... = (a ⊓ b) * ((a * b) * (a ⊓ b)⁻¹) : by rw [inv_inf_eq_sup_inv, sup_comm] ... = a * b : by rw [mul_comm, inv_mul_cancel_right] @@ -175,23 +166,18 @@ lemma one_le_neg (a : α) : 1 ≤ a⁻ := le_sup_right @[to_additive] -- pos_nonpos_iff lemma pos_le_one_iff {a : α} : a⁺ ≤ 1 ↔ a ≤ 1 := -by { rw [m_pos_part_def, sup_le_iff], simp, } +by rw [m_pos_part_def, sup_le_iff, and_iff_left le_rfl] @[to_additive] -- neg_nonpos_iff lemma neg_le_one_iff {a : α} : a⁻ ≤ 1 ↔ a⁻¹ ≤ 1 := -by { rw [m_neg_part_def, sup_le_iff], simp, } - -@[to_additive] -lemma pos_eq_one_iff {a : α} : a⁺ = 1 ↔ a ≤ 1 := -by { rw le_antisymm_iff, simp only [one_le_pos, and_true], exact pos_le_one_iff, } +by rw [m_neg_part_def, sup_le_iff, and_iff_left le_rfl] -@[to_additive] -lemma neg_eq_one_iff' {a : α} : a⁻ = 1 ↔ a⁻¹ ≤ 1 := -by { rw le_antisymm_iff, simp only [one_le_neg, and_true], rw neg_le_one_iff, } +@[to_additive] lemma pos_eq_one_iff {a : α} : a⁺ = 1 ↔ a ≤ 1 := sup_eq_right +@[to_additive] lemma neg_eq_one_iff' {a : α} : a⁻ = 1 ↔ a⁻¹ ≤ 1 := sup_eq_right @[to_additive] lemma neg_eq_one_iff [covariant_class α α has_mul.mul has_le.le] {a : α} : a⁻ = 1 ↔ 1 ≤ a := -by { rw le_antisymm_iff, simp only [one_le_neg, and_true], rw [neg_le_one_iff, inv_le_one'], } +by rw [le_antisymm_iff, neg_le_one_iff, inv_le_one', and_iff_left (one_le_neg _)] @[to_additive le_pos] lemma m_le_pos (a : α) : a ≤ a⁺ := le_sup_left @@ -207,18 +193,7 @@ lemma neg_eq_pos_inv (a : α) : a⁻ = (a⁻¹)⁺ := rfl -- a⁺ = (-a)⁻ @[to_additive] -lemma pos_eq_neg_inv (a : α) : a⁺ = (a⁻¹)⁻ := by simp [neg_eq_pos_inv] - --- We use this in Bourbaki A.VI.12 Prop 9 a) --- c + (a ⊓ b) = (c + a) ⊓ (c + b) -@[to_additive] -lemma mul_inf_eq_mul_inf_mul [covariant_class α α (*) (≤)] - (a b c : α) : c * (a ⊓ b) = (c * a) ⊓ (c * b) := -begin - refine le_antisymm (by simp) _, - rw [← mul_le_mul_iff_left c⁻¹, ← mul_assoc, inv_mul_self, one_mul, le_inf_iff], - simp, -end +lemma pos_eq_neg_inv (a : α) : a⁺ = (a⁻¹)⁻ := by rw [neg_eq_pos_inv, inv_inv] -- Bourbaki A.VI.12 Prop 9 a) -- a = a⁺ - a⁻ @@ -235,7 +210,7 @@ end -- a⁺ ⊓ a⁻ = 0 (`a⁺` and `a⁻` are co-prime, and, since they are positive, disjoint) @[to_additive] lemma pos_inf_neg_eq_one [covariant_class α α (*) (≤)] (a : α) : a⁺ ⊓ a⁻ = 1 := -by rw [←mul_right_inj (a⁻)⁻¹, mul_inf_eq_mul_inf_mul, mul_one, mul_left_inv, mul_comm, +by rw [←mul_right_inj (a⁻)⁻¹, mul_inf, mul_one, mul_left_inv, mul_comm, ← div_eq_mul_inv, pos_div_neg, neg_eq_inv_inf_one, inv_inv] -- Bourbaki A.VI.12 (with a and b swapped) @@ -252,7 +227,7 @@ calc a ⊔ b = (b * (a / b)) ⊔ (b * 1) : by rw [mul_one b, div_eq_mul_inv, mul lemma inf_eq_div_pos_div [covariant_class α α (*) (≤)] (a b : α) : a ⊓ b = a / (a / b)⁺ := calc a ⊓ b = (a * 1) ⊓ (a * (b / a)) : by { rw [mul_one a, div_eq_mul_inv, mul_comm b, mul_inv_cancel_left], } -... = a * (1 ⊓ (b / a)) : by rw ← mul_inf_eq_mul_inf_mul 1 (b / a) a +... = a * (1 ⊓ (b / a)) : by rw ← mul_inf 1 (b / a) a ... = a * ((b / a) ⊓ 1) : by rw inf_comm ... = a * ((a / b)⁻¹ ⊓ 1) : by { rw div_eq_mul_inv, nth_rewrite 0 ← inv_inv b, rw [← mul_inv, mul_comm b⁻¹, ← div_eq_mul_inv], } @@ -298,35 +273,22 @@ end lemma one_le_abs [covariant_class α α (*) (≤)] (a : α) : 1 ≤ |a| := by { rw ← m_pos_abs, exact one_le_pos _, } --- The proof from Bourbaki A.VI.12 Prop 9 d) -- |a| = a⁺ - a⁻ @[to_additive] lemma pos_mul_neg [covariant_class α α (*) (≤)] (a : α) : |a| = a⁺ * a⁻ := begin - refine le_antisymm _ _, - { refine sup_le _ _, - { nth_rewrite 0 ← mul_one a, - exact mul_le_mul' (m_le_pos a) (one_le_neg a) }, - { nth_rewrite 0 ← one_mul (a⁻¹), - exact mul_le_mul' (one_le_pos a) (inv_le_neg a) } }, - { rw [← inf_mul_sup, pos_inf_neg_eq_one, one_mul, ← m_pos_abs a], - apply sup_le, - { exact ((m_le_iff_pos_le_neg_ge _ _).mp (le_mabs a)).left, }, - { rw neg_eq_pos_inv, - exact ((m_le_iff_pos_le_neg_ge _ _).mp (inv_le_abs a)).left, }, } + rw [m_pos_part_def, sup_mul, one_mul, m_neg_part_def, mul_sup, mul_one, mul_inv_self, sup_assoc, + ←@sup_assoc _ _ a, sup_eq_right.2 le_sup_right], + exact (sup_eq_left.2 $ one_le_abs a).symm, end -- a ⊔ b - (a ⊓ b) = |b - a| @[to_additive] lemma sup_div_inf_eq_abs_div [covariant_class α α (*) (≤)] (a b : α) : - (a ⊔ b) / (a ⊓ b) = |b / a| := -begin - rw [sup_eq_mul_pos_div, inf_comm, inf_eq_div_pos_div, div_eq_mul_inv], - nth_rewrite 1 div_eq_mul_inv, - rw [mul_inv_rev, inv_inv, mul_comm, ← mul_assoc, inv_mul_cancel_right, pos_eq_neg_inv (a / b)], - nth_rewrite 1 div_eq_mul_inv, - rw [mul_inv_rev, ← div_eq_mul_inv, inv_inv, ← pos_mul_neg], -end + (a ⊔ b) / (a ⊓ b) = |b / a| := by +rw [sup_eq_mul_pos_div, inf_comm, inf_eq_div_pos_div, div_eq_mul_inv, div_eq_mul_inv b ((b / a)⁺), + mul_inv_rev, inv_inv, mul_comm, ← mul_assoc, inv_mul_cancel_right, pos_eq_neg_inv (a / b), + div_eq_mul_inv a b, mul_inv_rev, ← div_eq_mul_inv, inv_inv, ← pos_mul_neg] -- 2•(a ⊔ b) = a + b + |b - a| @[to_additive two_sup_eq_add_add_abs_sub] @@ -379,12 +341,9 @@ begin ((b ⊔ c ⊔ (a ⊔ c)) / ((b ⊔ c) ⊓ (a ⊔ c))) * |(a ⊓ c) / (b ⊓ c)| : by rw sup_div_inf_eq_abs_div ... = (b ⊔ c ⊔ (a ⊔ c)) / ((b ⊔ c) ⊓ (a ⊔ c)) * (((b ⊓ c) ⊔ (a ⊓ c)) / ((b ⊓ c) ⊓ (a ⊓ c))) : by rw sup_div_inf_eq_abs_div (b ⊓ c) (a ⊓ c) - ... = (b ⊔ a ⊔ c) / ((b ⊓ a) ⊔ c) * (((b ⊔ a) ⊓ c) / (b ⊓ a ⊓ c)) : by - { rw [← sup_inf_right, ← inf_sup_right, sup_assoc], - nth_rewrite 1 sup_comm, - rw [sup_right_idem, sup_assoc, inf_assoc], - nth_rewrite 3 inf_comm, - rw [inf_right_idem, inf_assoc], } + ... = (b ⊔ a ⊔ c) / ((b ⊓ a) ⊔ c) * (((b ⊔ a) ⊓ c) / (b ⊓ a ⊓ c)) : + by rw [← sup_inf_right, ← inf_sup_right, sup_assoc, @sup_comm _ _ c (a⊔c), sup_right_idem, + sup_assoc, inf_assoc, @inf_comm _ _ c (a⊓c), inf_right_idem, inf_assoc] ... = (b ⊔ a ⊔ c) * ((b ⊔ a) ⊓ c) /(((b ⊓ a) ⊔ c) * (b ⊓ a ⊓ c)) : by rw div_mul_div_comm ... = (b ⊔ a) * c / ((b ⊓ a) * c) : by rw [mul_comm, inf_mul_sup, mul_comm (b ⊓ a ⊔ c), inf_mul_sup] @@ -422,7 +381,7 @@ neg_eq_one_iff'.mpr h @[to_additive] -- neg_of_nonpos lemma neg_of_le_one [covariant_class α α (*) (≤)] (a : α) (h : a ≤ 1) : a⁻ = a⁻¹ := -by { refine neg_of_one_le_inv _ _, rw one_le_inv', exact h, } +sup_eq_left.2 $ one_le_inv'.2 h @[to_additive] -- neg_of_nonneg' lemma neg_of_one_le [covariant_class α α (*) (≤)] (a : α) (h : 1 ≤ a) : a⁻ = 1 := @@ -431,13 +390,7 @@ neg_eq_one_iff.mpr h -- 0 ≤ a implies |a| = a @[to_additive abs_of_nonneg] lemma mabs_of_one_le [covariant_class α α (*) (≤)] (a : α) (h : 1 ≤ a) : |a| = a := -begin - unfold has_abs.abs, - rw [sup_eq_mul_pos_div, div_eq_mul_inv, inv_inv, ← pow_two, inv_mul_eq_iff_eq_mul, - ← pow_two, pos_of_one_le], - rw pow_two, - apply one_le_mul h h, -end +sup_eq_left.2 $ left.inv_le_self h /-- The unary operation of taking the absolute value is idempotent. -/ @[simp, to_additive abs_abs "The unary operation of taking the absolute value is idempotent."] @@ -495,19 +448,15 @@ end @[to_additive] lemma abs_abs_div_abs_le [covariant_class α α (*) (≤)] (a b : α) : | |a| / |b| | ≤ |a / b| := begin - unfold has_abs.abs, - rw sup_le_iff, + rw [abs_eq_sup_inv, sup_le_iff], split, { apply div_le_iff_le_mul.2, convert mabs_mul_le (a/b) b, - { rw div_mul_cancel', }, - { rw div_mul_cancel', }, - { exact covariant_swap_mul_le_of_covariant_mul_le α, } }, - { rw [div_eq_mul_inv, mul_inv_rev, inv_inv, mul_inv_le_iff_le_mul, ← abs_eq_sup_inv (a / b), - abs_inv_comm], + rw div_mul_cancel', + exact covariant_swap_mul_le_of_covariant_mul_le α, }, + { rw [div_eq_mul_inv, mul_inv_rev, inv_inv, mul_inv_le_iff_le_mul, abs_inv_comm], convert mabs_mul_le (b/a) a, - { rw div_mul_cancel', }, - {rw div_mul_cancel', } }, + { rw div_mul_cancel', }, }, end end lattice_ordered_comm_group From 3c1368cac4abd5a5cbe44317ba7e87379d51ed88 Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Sat, 1 Apr 2023 12:40:11 +0000 Subject: [PATCH 0750/1291] fix(data/nat/squarefree): `norm_num` only supports `squarefree` on naturals (#18708) We have a `norm_num` extension for proving whether natural numbers are squarefree. This tactic does not work on other rings, like integers, but the check for this case was broken, meaning it returned a proof with a type error instead of skipping the goal. This PR changes the check to only fire the `norm_num` extension on natural numbers. I've included a small test that checks `norm_num` fails on integers, instead of returning a type-incorrect proof. We'd have to replace this test if `norm_num` gains support for `squarefree` on integers. --- src/data/nat/squarefree.lean | 3 ++- test/norm_num_ext.lean | 23 +++++++++++++++++++++++ 2 files changed, 25 insertions(+), 1 deletion(-) diff --git a/src/data/nat/squarefree.lean b/src/data/nat/squarefree.lean index f0ab22e538b91..08150e41ae27c 100644 --- a/src/data/nat/squarefree.lean +++ b/src/data/nat/squarefree.lean @@ -513,7 +513,8 @@ end /-- Evaluates the `squarefree` predicate on naturals. -/ @[norm_num] meta def eval_squarefree : expr → tactic (expr × expr) -| `(squarefree (%%e : ℕ)) := do +| `(@squarefree ℕ %%inst %%e) := do + is_def_eq inst `(nat.monoid), n ← e.to_nat, match n with | 0 := false_intro `(@not_squarefree_zero ℕ _ _) diff --git a/test/norm_num_ext.lean b/test/norm_num_ext.lean index 7009667fa041c..d33578ed6e431 100644 --- a/test/norm_num_ext.lean +++ b/test/norm_num_ext.lean @@ -243,6 +243,29 @@ example : squarefree 10 := by norm_num example : squarefree (2*3*5*17) := by norm_num example : ¬ squarefree (2*3*5*5*17) := by norm_num example : squarefree 251 := by norm_num +example : squarefree (3 : ℤ) := +begin + -- `norm_num` should fail on this example, instead of producing an incorrect proof. + success_if_fail { norm_num }, + exact irreducible.squarefree (prime.irreducible + (int.prime_iff_nat_abs_prime.mpr (by norm_num))) +end +example : @squarefree ℕ multiplicative.monoid 1 := +begin + -- `norm_num` should fail on this example, instead of producing an incorrect proof. + success_if_fail { norm_num }, + -- the statement was deliberately wacky, let's fix it + change squarefree (multiplicative.of_add 1 : multiplicative ℕ), + rintros x ⟨dx, hd⟩, + revert x dx, + rw multiplicative.of_add.surjective.forall₂, + intros x dx h, + simp_rw [←of_add_add, multiplicative.of_add.injective.eq_iff] at h, + cases x, + { simp [is_unit_one], exact is_unit_one }, + { simp only [nat.succ_add, nat.add_succ] at h, + cases h }, +end example : nat.fib 0 = 0 := by norm_num example : nat.fib 1 = 1 := by norm_num From 207c92594599a06e7c134f8d00a030a83e6c7259 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 1 Apr 2023 15:55:12 +0000 Subject: [PATCH 0751/1291] feat(combinatorics/additive/e_transform): e-transforms (#18683) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Define the general e-transform along with two special cases: * The Dyson e-transform, which turns `(s, t)` into `(s ∪ e • t, t ∩ e⁻¹ • s)`. * An pair of unnamed transforms, which turn `(s, t)` into `(s ∩ s • e, t ∪ e⁻¹ • t)` and `(s ∪ s • e, t ∩ e⁻¹ • t)`. --- src/combinatorics/additive/e_transform.lean | 159 ++++++++++++++++++++ 1 file changed, 159 insertions(+) create mode 100644 src/combinatorics/additive/e_transform.lean diff --git a/src/combinatorics/additive/e_transform.lean b/src/combinatorics/additive/e_transform.lean new file mode 100644 index 0000000000000..1f99c723f6a8a --- /dev/null +++ b/src/combinatorics/additive/e_transform.lean @@ -0,0 +1,159 @@ +/- +Copyright (c) 2023 Yaël Dillies. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies +-/ +import data.finset.pointwise + +/-! +# e-transforms + +e-transforms are a family of transformations of pairs of finite sets that aim to reduce the size +of the sumset while keeping some invariant the same. This file defines a few of them, to be used +as internals of other proofs. + +## Main declarations + +* `finset.mul_dyson_e_transform`: The Dyson e-transform. Replaces `(s, t)` by + `(s ∪ e • t, t ∩ e⁻¹ • s)`. The additive version preserves `|s ∩ [1, m]| + |t ∩ [1, m - e]|`. +* `finset.mul_e_transform_left`/`finset.mul_e_transform_right`: Replace `(s, t)` by + `(s ∩ s • e, t ∪ e⁻¹ • t)` and `(s ∪ s • e, t ∩ e⁻¹ • t)`. Preserve (together) the sum of + the cardinalities (see `finset.mul_e_transform.card`). In particular, one of the two transforms + increases the sum of the cardinalities and the other one decreases it. See + `le_or_lt_of_add_le_add` and around. + +## TODO + +Prove the invariance property of the Dyson e-transform. +-/ + +open mul_opposite +open_locale pointwise + +variables {α : Type*} [decidable_eq α] + +namespace finset + +/-! ### Dyson e-transform -/ + +section comm_group +variables [comm_group α] (e : α) (x : finset α × finset α) + +/-- The **Dyson e-transform**. Turns `(s, t)` into `(s ∪ e • t, t ∩ e⁻¹ • s)`. This reduces the +product of the two sets. -/ +@[to_additive "The **Dyson e-transform**. Turns `(s, t)` into `(s ∪ e +ᵥ t, t ∩ -e +ᵥ s)`. This +reduces the sum of the two sets.", simps] +def mul_dyson_e_transform : finset α × finset α := (x.1 ∪ e • x.2, x.2 ∩ e⁻¹ • x.1) + +@[to_additive] lemma mul_dyson_e_transform.subset : + (mul_dyson_e_transform e x).1 * (mul_dyson_e_transform e x).2 ⊆ x.1 * x.2 := +begin + refine union_mul_inter_subset_union.trans (union_subset subset.rfl _), + rw [mul_smul_comm, smul_mul_assoc, inv_smul_smul, mul_comm], + refl, +end + +@[to_additive] lemma mul_dyson_e_transform.card : + (mul_dyson_e_transform e x).1.card + (mul_dyson_e_transform e x).2.card = x.1.card + x.2.card := +begin + dsimp, + rw [←card_smul_finset e (_ ∩ _), smul_finset_inter, smul_inv_smul, inter_comm, + card_union_add_card_inter, card_smul_finset], +end + +@[simp, to_additive] lemma mul_dyson_e_transform_idem : + mul_dyson_e_transform e (mul_dyson_e_transform e x) = mul_dyson_e_transform e x := +begin + ext : 1; dsimp, + { rw [smul_finset_inter, smul_inv_smul, inter_comm, union_eq_left_iff_subset], + exact inter_subset_union }, + { rw [smul_finset_union, inv_smul_smul, union_comm, inter_eq_left_iff_subset], + exact inter_subset_union } +end + +variables {e x} + +@[to_additive] lemma mul_dyson_e_transform.smul_finset_snd_subset_fst : + e • (mul_dyson_e_transform e x).2 ⊆ (mul_dyson_e_transform e x).1 := +by { dsimp, rw [smul_finset_inter, smul_inv_smul, inter_comm], exact inter_subset_union } + +end comm_group + +/-! +### Two unnamed e-transforms + +The following two transforms both reduce the product/sum of the two sets. Further, one of them must +decrease the sum of the size of the sets (and then the other increases it). + +This pair of transforms doesn't seem to be named in the literature. It is used by Sanders in his +bound on Roth numbers, and by DeVos in his proof of Cauchy-Davenport. +-/ + +section group +variables [group α] (e : α) (x : finset α × finset α) + +/-- An **e-transform**. Turns `(s, t)` into `(s ∩ s • e, t ∪ e⁻¹ • t)`. This reduces the +product of the two sets. -/ +@[to_additive "An **e-transform**. Turns `(s, t)` into `(s ∩ s +ᵥ e, t ∪ -e +ᵥ t)`. This +reduces the sum of the two sets.", simps] +def mul_e_transform_left : finset α × finset α := (x.1 ∩ op e • x.1, x.2 ∪ e⁻¹ • x.2) + +/-- An **e-transform**. Turns `(s, t)` into `(s ∪ s • e, t ∩ e⁻¹ • t)`. This reduces the product of +the two sets. -/ +@[to_additive "An **e-transform**. Turns `(s, t)` into `(s ∪ s +ᵥ e, t ∩ -e +ᵥ t)`. This reduces the +sum of the two sets.", simps] +def mul_e_transform_right : finset α × finset α := (x.1 ∪ op e • x.1, x.2 ∩ e⁻¹ • x.2) + +@[simp, to_additive] lemma mul_e_transform_left_one : mul_e_transform_left 1 x = x := +by simp [mul_e_transform_left] +@[simp, to_additive] lemma mul_e_transform_right_one : mul_e_transform_right 1 x = x := +by simp [mul_e_transform_right] + +@[to_additive] lemma mul_e_transform_left.fst_mul_snd_subset : + (mul_e_transform_left e x).1 * (mul_e_transform_left e x).2 ⊆ x.1 * x.2 := +begin + refine inter_mul_union_subset_union.trans (union_subset subset.rfl _), + rw [op_smul_finset_mul_eq_mul_smul_finset, smul_inv_smul], + refl, +end + +@[to_additive] lemma mul_e_transform_right.fst_mul_snd_subset : + (mul_e_transform_right e x).1 * (mul_e_transform_right e x).2 ⊆ x.1 * x.2 := +begin + refine union_mul_inter_subset_union.trans (union_subset subset.rfl _), + rw [op_smul_finset_mul_eq_mul_smul_finset, smul_inv_smul], + refl, +end + +@[to_additive] lemma mul_e_transform_left.card : + (mul_e_transform_left e x).1.card + (mul_e_transform_right e x).1.card = 2 * x.1.card := +(card_inter_add_card_union _ _).trans $ by rw [card_smul_finset, two_mul] + +@[to_additive] lemma mul_e_transform_right.card : + (mul_e_transform_left e x).2.card + (mul_e_transform_right e x).2.card = 2 * x.2.card := +(card_union_add_card_inter _ _).trans $ by rw [card_smul_finset, two_mul] + +/-- This statement is meant to be combined with `le_or_lt_of_add_le_add` and similar lemmas. -/ +@[to_additive add_e_transform.card "This statement is meant to be combined with +`le_or_lt_of_add_le_add` and similar lemmas."] +protected lemma mul_e_transform.card : + (mul_e_transform_left e x).1.card + (mul_e_transform_left e x).2.card + + ((mul_e_transform_right e x).1.card + (mul_e_transform_right e x).2.card) + = x.1.card + x.2.card + (x.1.card + x.2.card) := +by rw [add_add_add_comm, mul_e_transform_left.card, mul_e_transform_right.card, ←mul_add, two_mul] + +end group + +section comm_group +variables [comm_group α] (e : α) (x : finset α × finset α) + +@[simp, to_additive] lemma mul_e_transform_left_inv : + mul_e_transform_left e⁻¹ x = (mul_e_transform_right e x.swap).swap := +by simp [-op_inv, op_smul_eq_smul, mul_e_transform_left, mul_e_transform_right] + +@[simp, to_additive] lemma mul_e_transform_right_inv : + mul_e_transform_right e⁻¹ x = (mul_e_transform_left e x.swap).swap := +by simp [-op_inv, op_smul_eq_smul, mul_e_transform_left, mul_e_transform_right] + +end comm_group +end finset From 88fcb83fe7996142dfcfe7368d31304a9adc874a Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 1 Apr 2023 17:10:53 +0000 Subject: [PATCH 0752/1291] feat(measure_theory/measurable_space): Cast of natural is measurable (#18676) A few simple lemmas --- src/measure_theory/measurable_space.lean | 6 ++++++ src/measure_theory/measure/measure_space.lean | 4 ++++ 2 files changed, 10 insertions(+) diff --git a/src/measure_theory/measurable_space.lean b/src/measure_theory/measurable_space.lean index d9c69d5c5e234..01bf4ed280315 100644 --- a/src/measure_theory/measurable_space.lean +++ b/src/measure_theory/measurable_space.lean @@ -209,6 +209,12 @@ begin { convert measurable_const, exact funext (λ x, hf x h.some) } end +@[measurability] lemma measurable_nat_cast [has_nat_cast α] (n : ℕ) : measurable (n : β → α) := +@measurable_const α _ _ _ n + +@[measurability] lemma measurable_int_cast [has_int_cast α] (n : ℤ) : measurable (n : β → α) := +@measurable_const α _ _ _ n + lemma measurable_of_finite [finite α] [measurable_singleton_class α] (f : α → β) : measurable f := λ s hs, (f ⁻¹' s).to_finite.measurable_set diff --git a/src/measure_theory/measure/measure_space.lean b/src/measure_theory/measure/measure_space.lean index 02af7350e7fa0..1fbaec96e712a 100644 --- a/src/measure_theory/measure/measure_space.lean +++ b/src/measure_theory/measure/measure_space.lean @@ -946,6 +946,10 @@ lemma nonpos_iff_eq_zero' : μ ≤ 0 ↔ μ = 0 := @[simp] lemma measure_univ_eq_zero : μ univ = 0 ↔ μ = 0 := ⟨λ h, bot_unique $ λ s hs, trans_rel_left (≤) (measure_mono (subset_univ s)) h, λ h, h.symm ▸ rfl⟩ +lemma measure_univ_ne_zero : μ univ ≠ 0 ↔ μ ≠ 0 := measure_univ_eq_zero.not + +@[simp] lemma measure_univ_pos : 0 < μ univ ↔ μ ≠ 0 := pos_iff_ne_zero.trans measure_univ_ne_zero + /-! ### Pushforward and pullback -/ /-- Lift a linear map between `outer_measure` spaces such that for each measure `μ` every measurable From 6b60020790e39e77bfd633ba3d5562ff82e52c79 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sun, 2 Apr 2023 09:13:02 +0000 Subject: [PATCH 0753/1291] chore(group_theory/subgroup/basic): Protect `subgroup.subtype` (#18712) and `subgroup_class.subtype`. This was breaking the `subtype` notation for me in #18684. The longer term fix would be to add a `_root_` in the notation declaration. --- src/group_theory/subgroup/basic.lean | 14 +++++++------- 1 file changed, 7 insertions(+), 7 deletions(-) diff --git a/src/group_theory/subgroup/basic.lean b/src/group_theory/subgroup/basic.lean index f93c310ab51bb..b18053c47065b 100644 --- a/src/group_theory/subgroup/basic.lean +++ b/src/group_theory/subgroup/basic.lean @@ -221,18 +221,18 @@ include hSG /-- The natural group hom from a subgroup of group `G` to `G`. -/ @[to_additive "The natural group hom from an additive subgroup of `add_group` `G` to `G`."] -def subtype : H →* G := ⟨coe, rfl, λ _ _, rfl⟩ +protected def subtype : H →* G := ⟨coe, rfl, λ _ _, rfl⟩ -@[simp, to_additive] theorem coe_subtype : (subtype H : H → G) = coe := rfl +@[simp, to_additive] theorem coe_subtype : (subgroup_class.subtype H : H → G) = coe := rfl variables {H} @[simp, norm_cast, to_additive coe_smul] lemma coe_pow (x : H) (n : ℕ) : ((x ^ n : H) : G) = x ^ n := -(subtype H : H →* G).map_pow _ _ +(subgroup_class.subtype H : H →* G).map_pow _ _ @[simp, norm_cast, to_additive] lemma coe_zpow (x : H) (n : ℤ) : ((x ^ n : H) : G) = x ^ n := -(subtype H : H →* G).map_zpow _ _ +(subgroup_class.subtype H : H →* G).map_zpow _ _ /-- The inclusion homomorphism from a subgroup `H` contained in `K` to `K`. -/ @[to_additive "The inclusion homomorphism from a additive subgroup `H` contained in `K` to `K`."] @@ -257,7 +257,7 @@ by { cases a, simp only [inclusion, set_like.coe_mk, monoid_hom.mk'_apply] } @[simp, to_additive] lemma subtype_comp_inclusion {H K : S} (hH : H ≤ K) : - (subtype K).comp (inclusion hH) = subtype H := + (subgroup_class.subtype K).comp (inclusion hH) = subgroup_class.subtype H := by { ext, simp only [monoid_hom.comp_apply, coe_subtype, coe_inclusion] } end subgroup_class @@ -538,11 +538,11 @@ subtype.coe_injective.linear_ordered_comm_group _ /-- The natural group hom from a subgroup of group `G` to `G`. -/ @[to_additive "The natural group hom from an `add_subgroup` of `add_group` `G` to `G`."] -def subtype : H →* G := ⟨coe, rfl, λ _ _, rfl⟩ +protected def subtype : H →* G := ⟨coe, rfl, λ _ _, rfl⟩ @[simp, to_additive] theorem coe_subtype : ⇑H.subtype = coe := rfl -@[to_additive] lemma subtype_injective : function.injective (subtype H) := subtype.coe_injective +@[to_additive] lemma subtype_injective : injective (subgroup.subtype H) := subtype.coe_injective /-- The inclusion homomorphism from a subgroup `H` contained in `K` to `K`. -/ @[to_additive "The inclusion homomorphism from a additive subgroup `H` contained in `K` to `K`."] From ea050b44c0f9aba9d16a948c7cc7d2e7c8493567 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Sun, 2 Apr 2023 12:33:21 +0000 Subject: [PATCH 0754/1291] feat(set_theory/cardinal/basic): clean up instances (#18714) We add a missing `linear_ordered_comm_monoid_with_zero` instance and reorder others. --- src/set_theory/cardinal/basic.lean | 33 +++++++++++++++++------------- 1 file changed, 19 insertions(+), 14 deletions(-) diff --git a/src/set_theory/cardinal/basic.lean b/src/set_theory/cardinal/basic.lean index e64e91c84d436..5022c5826899b 100644 --- a/src/set_theory/cardinal/basic.lean +++ b/src/set_theory/cardinal/basic.lean @@ -190,11 +190,13 @@ instance : has_le cardinal.{u} := ⟨λ q₁ q₂, quotient.lift_on₂ q₁ q₂ (λ α β, nonempty $ α ↪ β) $ λ α β γ δ ⟨e₁⟩ ⟨e₂⟩, propext ⟨λ ⟨e⟩, ⟨e.congr e₁ e₂⟩, λ ⟨e⟩, ⟨e.congr e₁.symm e₂.symm⟩⟩⟩ -instance : partial_order cardinal.{u} := -{ le := (≤), - le_refl := by rintros ⟨α⟩; exact ⟨embedding.refl _⟩, - le_trans := by rintros ⟨α⟩ ⟨β⟩ ⟨γ⟩ ⟨e₁⟩ ⟨e₂⟩; exact ⟨e₁.trans e₂⟩, - le_antisymm := by { rintros ⟨α⟩ ⟨β⟩ ⟨e₁⟩ ⟨e₂⟩, exact quotient.sound (e₁.antisymm e₂) } } +instance : linear_order cardinal.{u} := +{ le := (≤), + le_refl := by rintros ⟨α⟩; exact ⟨embedding.refl _⟩, + le_trans := by rintros ⟨α⟩ ⟨β⟩ ⟨γ⟩ ⟨e₁⟩ ⟨e₂⟩; exact ⟨e₁.trans e₂⟩, + le_antisymm := by { rintros ⟨α⟩ ⟨β⟩ ⟨e₁⟩ ⟨e₂⟩, exact quotient.sound (e₁.antisymm e₂) }, + le_total := by { rintros ⟨α⟩ ⟨β⟩, apply embedding.total }, + decidable_le := classical.dec_rel _ } theorem le_def (α β : Type u) : #α ≤ #β ↔ nonempty (α ↪ β) := iff.rfl @@ -472,7 +474,17 @@ instance : canonically_ordered_comm_semiring cardinal.{u} := le_self_add := λ a b, (add_zero a).ge.trans $ add_le_add_left (cardinal.zero_le _) _, eq_zero_or_eq_zero_of_mul_eq_zero := λ a b, induction_on₂ a b $ λ α β, by simpa only [mul_def, mk_eq_zero_iff, is_empty_prod] using id, - ..cardinal.comm_semiring, ..cardinal.partial_order } + ..cardinal.comm_semiring, ..cardinal.linear_order } + +instance : canonically_linear_ordered_add_monoid cardinal.{u} := +{ ..cardinal.canonically_ordered_comm_semiring, + ..cardinal.linear_order } + +instance : linear_ordered_comm_monoid_with_zero cardinal.{u} := +{ mul_le_mul_left := @mul_le_mul_left' _ _ _ _, + zero_le_one := zero_le _, + ..cardinal.comm_semiring, + ..cardinal.linear_order } lemma zero_power_le (c : cardinal.{u}) : (0 : cardinal.{u}) ^ c ≤ 1 := by { by_cases h : c = 0, rw [h, power_zero], rw [zero_power h], apply zero_le } @@ -500,13 +512,7 @@ begin end instance : no_max_order cardinal.{u} := -{ exists_gt := λ a, ⟨_, cantor a⟩, ..cardinal.partial_order } - -instance : canonically_linear_ordered_add_monoid cardinal.{u} := -{ le_total := by { rintros ⟨α⟩ ⟨β⟩, apply embedding.total }, - decidable_le := classical.dec_rel _, - ..(infer_instance : canonically_ordered_add_monoid cardinal), - ..cardinal.partial_order } +{ exists_gt := λ a, ⟨_, cantor a⟩ } -- short-circuit type class inference instance : distrib_lattice cardinal.{u} := by apply_instance @@ -541,7 +547,6 @@ protected theorem lt_wf : @well_founded cardinal.{u} (<) := end⟩ instance : has_well_founded cardinal.{u} := ⟨(<), cardinal.lt_wf⟩ -instance : well_founded_lt cardinal.{u} := ⟨cardinal.lt_wf⟩ instance wo : @is_well_order cardinal.{u} (<) := { } instance : conditionally_complete_linear_order_bot cardinal := From 5f6e827d81dfbeb6151d7016586ceeb0099b9655 Mon Sep 17 00:00:00 2001 From: Oliver Nash Date: Sun, 2 Apr 2023 14:57:21 +0000 Subject: [PATCH 0755/1291] chore(measure_theory/measure/doubling): rename `is_doubling_measure` to `is_unif_loc_doubling_measure` (#18709) With thanks to Jireh for highlighting our non-standard terminology. See also [Zulip conversation](https://leanprover.zulipchat.com/#narrow/stream/116395-maths/topic/definition.20of.20doubling.20measures). --- src/analysis/calculus/monotone.lean | 2 +- src/dynamics/ergodic/add_circle.lean | 3 +- src/measure_theory/covering/besicovitch.lean | 2 +- .../covering/density_theorem.lean | 33 +++++++------ .../covering/liminf_limsup.lean | 31 ++++++------ src/measure_theory/covering/one_dim.lean | 2 +- src/measure_theory/integral/periodic.lean | 2 +- src/measure_theory/measure/doubling.lean | 49 ++++++++++--------- src/measure_theory/measure/haar_lebesgue.lean | 3 +- 9 files changed, 68 insertions(+), 59 deletions(-) diff --git a/src/analysis/calculus/monotone.lean b/src/analysis/calculus/monotone.lean index cc6205136f304..cabfe29cebc40 100644 --- a/src/analysis/calculus/monotone.lean +++ b/src/analysis/calculus/monotone.lean @@ -30,7 +30,7 @@ limit of `(f y - f x) / (y - x)` by a lower and upper approximation argument fro behavior of `μ [x, y]`. -/ -open set filter function metric measure_theory measure_theory.measure is_doubling_measure +open set filter function metric measure_theory measure_theory.measure is_unif_loc_doubling_measure open_locale topology /-- If `(f y - f x) / (y - x)` converges to a limit as `y` tends to `x`, then the same goes if diff --git a/src/dynamics/ergodic/add_circle.lean b/src/dynamics/ergodic/add_circle.lean index 7db687edc7e6c..18955e864bfe8 100644 --- a/src/dynamics/ergodic/add_circle.lean +++ b/src/dynamics/ergodic/add_circle.lean @@ -61,7 +61,8 @@ begin obtain ⟨d, -, hd⟩ : ∃ d, d ∈ s ∧ ∀ {ι'} {l : filter ι'} (w : ι' → add_circle T) (δ : ι' → ℝ), tendsto δ l (𝓝[>] 0) → (∀ᶠ j in l, d ∈ closed_ball (w j) (1 * δ j)) → tendsto (λ j, μ (s ∩ closed_ball (w j) (δ j)) / μ (closed_ball (w j) (δ j))) l (𝓝 1) := - exists_mem_of_measure_ne_zero_of_ae h (is_doubling_measure.ae_tendsto_measure_inter_div μ s 1), + exists_mem_of_measure_ne_zero_of_ae h + (is_unif_loc_doubling_measure.ae_tendsto_measure_inter_div μ s 1), let I : ι → set (add_circle T) := λ j, closed_ball d (T / (2 * ↑(n j))), replace hd : tendsto (λ j, μ (s ∩ I j) / μ (I j)) l (𝓝 1), { let δ : ι → ℝ := λ j, T / (2 * ↑(n j)), diff --git a/src/measure_theory/covering/besicovitch.lean b/src/measure_theory/covering/besicovitch.lean index 775b5ca0256da..9776ab4c2b7f3 100644 --- a/src/measure_theory/covering/besicovitch.lean +++ b/src/measure_theory/covering/besicovitch.lean @@ -1140,7 +1140,7 @@ to `1` when `r` tends to `0`, for almost every `x` in `s`. This shows that almost every point of `s` is a Lebesgue density point for `s`. A stronger version holds for measurable sets, see `ae_tendsto_measure_inter_div_of_measurable_set`. -See also `is_doubling_measure.ae_tendsto_measure_inter_div`. -/ +See also `is_unif_loc_doubling_measure.ae_tendsto_measure_inter_div`. -/ lemma ae_tendsto_measure_inter_div (μ : measure β) [is_locally_finite_measure μ] (s : set β) : ∀ᵐ x ∂(μ.restrict s), tendsto (λ r, μ (s ∩ (closed_ball x r)) / μ (closed_ball x r)) (𝓝[>] 0) (𝓝 1) := diff --git a/src/measure_theory/covering/density_theorem.lean b/src/measure_theory/covering/density_theorem.lean index 0cc20e2988eb0..b94da4ad6e87f 100644 --- a/src/measure_theory/covering/density_theorem.lean +++ b/src/measure_theory/covering/density_theorem.lean @@ -8,19 +8,21 @@ import measure_theory.covering.vitali import measure_theory.covering.differentiation /-! -# Doubling measures and Lebesgue's density theorem +# Uniformly locally doubling measures and Lebesgue's density theorem Lebesgue's density theorem states that given a set `S` in a sigma compact metric space with -locally-finite doubling measure `μ` then for almost all points `x` in `S`, for any sequence of -closed balls `B₀, B₁, B₂, ...` containing `x`, the limit `μ (S ∩ Bⱼ) / μ (Bⱼ) → 1` as `j → ∞`. +locally-finite uniformly locally doubling measure `μ` then for almost all points `x` in `S`, for any +sequence of closed balls `B₀, B₁, B₂, ...` containing `x`, the limit `μ (S ∩ Bⱼ) / μ (Bⱼ) → 1` as +`j → ∞`. -In this file we combine general results about existence of Vitali families for doubling measures -with results about differentiation along a Vitali family to obtain an explicit form of Lebesgue's -density theorem. +In this file we combine general results about existence of Vitali families for uniformly locally +doubling measures with results about differentiation along a Vitali family to obtain an explicit +form of Lebesgue's density theorem. ## Main results - * `is_doubling_measure.ae_tendsto_measure_inter_div`: a version of Lebesgue's density theorem for - sequences of balls converging on a point but whose centres are not required to be fixed. + * `is_unif_loc_doubling_measure.ae_tendsto_measure_inter_div`: a version of Lebesgue's density + theorem for sequences of balls converging on a point but whose centres are not required to be + fixed. -/ @@ -29,17 +31,18 @@ noncomputable theory open set filter metric measure_theory topological_space open_locale nnreal topology -namespace is_doubling_measure +namespace is_unif_loc_doubling_measure -variables {α : Type*} [metric_space α] [measurable_space α] (μ : measure α) [is_doubling_measure μ] +variables {α : Type*} [metric_space α] [measurable_space α] + (μ : measure α) [is_unif_loc_doubling_measure μ] section variables [second_countable_topology α] [borel_space α] [is_locally_finite_measure μ] open_locale topology -/-- A Vitali family in a space with a doubling measure, designed so that the sets at `x` contain -all `closed_ball y r` when `dist x y ≤ K * r`. -/ +/-- A Vitali family in a space with a uniformly locally doubling measure, designed so that the sets +at `x` contain all `closed_ball y r` when `dist x y ≤ K * r`. -/ @[irreducible] def vitali_family (K : ℝ) : vitali_family μ := begin /- the Vitali covering theorem gives a family that works well at small scales, thanks to the @@ -59,8 +62,8 @@ begin (R / 4) (by linarith), end -/-- In the Vitali family `is_doubling_measure.vitali_family K`, the sets based at `x` contain all -balls `closed_ball y r` when `dist x y ≤ K * r`. -/ +/-- In the Vitali family `is_unif_loc_doubling_measure.vitali_family K`, the sets based at `x` +contain all balls `closed_ball y r` when `dist x y ≤ K * r`. -/ lemma closed_ball_mem_vitali_family_of_dist_le_mul {K : ℝ} {x y : α} {r : ℝ} (h : dist x y ≤ K * r) (rpos : 0 < r) : closed_ball y r ∈ (vitali_family μ K).sets_at x := @@ -171,4 +174,4 @@ using hx.comp (tendsto_closed_ball_filter_at μ _ _ δlim xmem) end applications -end is_doubling_measure +end is_unif_loc_doubling_measure diff --git a/src/measure_theory/covering/liminf_limsup.lean b/src/measure_theory/covering/liminf_limsup.lean index a53637132c1a2..b32a298afbbd7 100644 --- a/src/measure_theory/covering/liminf_limsup.lean +++ b/src/measure_theory/covering/liminf_limsup.lean @@ -6,17 +6,17 @@ Authors: Oliver Nash import measure_theory.covering.density_theorem /-! -# Liminf, limsup, and doubling measures. +# Liminf, limsup, and uniformly locally doubling measures. This file is a place to collect lemmas about liminf and limsup for subsets of a metric space -carrying a doubling measure. +carrying a uniformly locally doubling measure. ## Main results: * `blimsup_cthickening_mul_ae_eq`: the limsup of the closed thickening of a sequence of subsets - of a metric space is unchanged almost everywhere for a doubling measure if the sequence of - distances is multiplied by a positive scale factor. This is a generalisation of a result of - Cassels, appearing as Lemma 9 on page 217 of + of a metric space is unchanged almost everywhere for a uniformly locally doubling measure if the + sequence of distances is multiplied by a positive scale factor. This is a generalisation of a + result of Cassels, appearing as Lemma 9 on page 217 of [J.W.S. Cassels, *Some metrical theorems in Diophantine approximation. I*](cassels1950). * `blimsup_thickening_mul_ae_eq`: a variant of `blimsup_cthickening_mul_ae_eq` for thickenings rather than closed thickenings. @@ -28,7 +28,7 @@ open_locale nnreal ennreal topology variables {α : Type*} [metric_space α] [second_countable_topology α] [measurable_space α] [borel_space α] -variables (μ : measure α) [is_locally_finite_measure μ] [is_doubling_measure μ] +variables (μ : measure α) [is_locally_finite_measure μ] [is_unif_loc_doubling_measure μ] /-- This is really an auxiliary result en route to `blimsup_cthickening_ae_le_of_eventually_mul_le` (which is itself an auxiliary result en route to `blimsup_cthickening_mul_ae_eq`). @@ -56,7 +56,8 @@ begin We obtain our contradiction by showing that there exists `η < 1` such that `μ (W ∩ (B j)) / μ (B j) ≤ η` for sufficiently large `j`. In fact we claim that `η = 1 - C⁻¹` - is such a value where `C` is the scaling constant of `M⁻¹` for the doubling measure `μ`. + is such a value where `C` is the scaling constant of `M⁻¹` for the uniformly locally doubling + measure `μ`. To prove the claim, let `b j = closed_ball (w j) (M * r₁ (f j))` and for given `j` consider the sets `b j` and `W ∩ (B j)`. These are both subsets of `B j` and are disjoint for large enough `j` @@ -76,7 +77,7 @@ begin tendsto δ l (𝓝[>] 0) → (∀ᶠ j in l, d ∈ closed_ball (w j) (2 * δ j)) → tendsto (λ j, μ (W ∩ closed_ball (w j) (δ j)) / μ (closed_ball (w j) (δ j))) l (𝓝 1) := measure.exists_mem_of_measure_ne_zero_of_ae contra - (is_doubling_measure.ae_tendsto_measure_inter_div μ W 2), + (is_unif_loc_doubling_measure.ae_tendsto_measure_inter_div μ W 2), replace hd : d ∈ blimsup Y₁ at_top p := ((mem_diff _).mp hd).1, obtain ⟨f : ℕ → ℕ, hf⟩ := exists_forall_mem_of_has_basis_mem_blimsup' at_top_basis hd, simp only [forall_and_distrib] at hf, @@ -96,9 +97,9 @@ begin { exact mem_Union₂.mp (cthickening_subset_Union_closed_ball_of_lt (s (f j)) (by positivity) (lt_two_mul_self hrp') (hf₀ j)), }, }, choose w hw hw' using hf₀, - let C := is_doubling_measure.scaling_constant_of μ M⁻¹, + let C := is_unif_loc_doubling_measure.scaling_constant_of μ M⁻¹, have hC : 0 < C := - lt_of_lt_of_le zero_lt_one (is_doubling_measure.one_le_scaling_constant_of μ M⁻¹), + lt_of_lt_of_le zero_lt_one (is_unif_loc_doubling_measure.one_le_scaling_constant_of μ M⁻¹), suffices : ∃ (η < (1 : ℝ≥0)), ∀ᶠ j in at_top, μ (W ∩ closed_ball (w j) (r₁ (f j))) / μ (closed_ball (w j) (r₁ (f j))) ≤ η, { obtain ⟨η, hη, hη'⟩ := this, @@ -123,7 +124,7 @@ begin simp only [mem_Union, exists_prop], exact ⟨f j, ⟨hf₁ j, hj.le.trans (hf₂ j)⟩, ha⟩, }, have h₄ : ∀ᶠ j in at_top, μ (B j) ≤ C * μ (b j) := - (hr.eventually (is_doubling_measure.eventually_measure_le_scaling_constant_mul' + (hr.eventually (is_unif_loc_doubling_measure.eventually_measure_le_scaling_constant_mul' μ M hM)).mono (λ j hj, hj (w j)), refine (h₃.and h₄).mono (λ j hj₀, _), change μ (W ∩ B j) / μ (B j) ≤ ↑(1 - C⁻¹), @@ -175,8 +176,8 @@ end /-- Given a sequence of subsets `sᵢ` of a metric space, together with a sequence of radii `rᵢ` such that `rᵢ → 0`, the set of points which belong to infinitely many of the closed -`rᵢ`-thickenings of `sᵢ` is unchanged almost everywhere for a doubling measure if the `rᵢ` are all -scaled by a positive constant. +`rᵢ`-thickenings of `sᵢ` is unchanged almost everywhere for a uniformly locally doubling measure if +the `rᵢ` are all scaled by a positive constant. This lemma is a generalisation of Lemma 9 appearing on page 217 of [J.W.S. Cassels, *Some metrical theorems in Diophantine approximation. I*](cassels1950). @@ -256,8 +257,8 @@ end /-- Given a sequence of subsets `sᵢ` of a metric space, together with a sequence of radii `rᵢ` such that `rᵢ → 0`, the set of points which belong to infinitely many of the -`rᵢ`-thickenings of `sᵢ` is unchanged almost everywhere for a doubling measure if the `rᵢ` are all -scaled by a positive constant. +`rᵢ`-thickenings of `sᵢ` is unchanged almost everywhere for a uniformly locally doubling measure if +the `rᵢ` are all scaled by a positive constant. This lemma is a generalisation of Lemma 9 appearing on page 217 of [J.W.S. Cassels, *Some metrical theorems in Diophantine approximation. I*](cassels1950). diff --git a/src/measure_theory/covering/one_dim.lean b/src/measure_theory/covering/one_dim.lean index c1a8cb11bb377..1df64bd0d094b 100644 --- a/src/measure_theory/covering/one_dim.lean +++ b/src/measure_theory/covering/one_dim.lean @@ -14,7 +14,7 @@ in `density_theorems.lean`. In this file, we expand the API for this theory in o by showing that intervals belong to the relevant Vitali family. -/ -open set measure_theory is_doubling_measure filter +open set measure_theory is_unif_loc_doubling_measure filter open_locale topology namespace real diff --git a/src/measure_theory/integral/periodic.lean b/src/measure_theory/integral/periodic.lean index f1f77629745cc..c93969a6b5810 100644 --- a/src/measure_theory/integral/periodic.lean +++ b/src/measure_theory/integral/periodic.lean @@ -107,7 +107,7 @@ begin { simp [hε, min_eq_left (by linarith : T ≤ 2 * ε)], }, end -instance : is_doubling_measure (volume : measure (add_circle T)) := +instance : is_unif_loc_doubling_measure (volume : measure (add_circle T)) := begin refine ⟨⟨real.to_nnreal 2, filter.eventually_of_forall $ λ ε x, _⟩⟩, simp only [volume_closed_ball], diff --git a/src/measure_theory/measure/doubling.lean b/src/measure_theory/measure/doubling.lean index 26608590ed104..1ca5de2930d7d 100644 --- a/src/measure_theory/measure/doubling.lean +++ b/src/measure_theory/measure/doubling.lean @@ -7,19 +7,20 @@ import analysis.special_functions.log.base import measure_theory.measure.measure_space_def /-! -# Doubling measures +# Uniformly locally doubling measures -A doubling measure `μ` on a metric space is a measure for which there exists a constant `C` such -that for all sufficiently small radii `ε`, and for any centre, the measure of a ball of radius -`2 * ε` is bounded by `C` times the measure of the concentric ball of radius `ε`. +A uniformly locally doubling measure `μ` on a metric space is a measure for which there exists a +constant `C` such that for all sufficiently small radii `ε`, and for any centre, the measure of a +ball of radius `2 * ε` is bounded by `C` times the measure of the concentric ball of radius `ε`. -This file records basic files on doubling measures. +This file records basic facts about uniformly locally doubling measures. ## Main definitions - * `is_doubling_measure`: the definition of a doubling measure (as a typeclass). - * `is_doubling_measure.doubling_constant`: a function yielding the doubling constant `C` appearing - in the definition of a doubling measure. + * `is_unif_loc_doubling_measure`: the definition of a uniformly locally doubling measure (as a + typeclass). + * `is_unif_loc_doubling_measure.doubling_constant`: a function yielding the doubling constant `C` + appearing in the definition of a uniformly locally doubling measure. -/ noncomputable theory @@ -27,26 +28,28 @@ noncomputable theory open set filter metric measure_theory topological_space open_locale ennreal nnreal topology -/-- A measure `μ` is said to be a doubling measure if there exists a constant `C` such that for -all sufficiently small radii `ε`, and for any centre, the measure of a ball of radius `2 * ε` is -bounded by `C` times the measure of the concentric ball of radius `ε`. +/-- A measure `μ` is said to be a uniformly locally doubling measure if there exists a constant `C` +such that for all sufficiently small radii `ε`, and for any centre, the measure of a ball of radius +`2 * ε` is bounded by `C` times the measure of the concentric ball of radius `ε`. Note: it is important that this definition makes a demand only for sufficiently small `ε`. For -example we want hyperbolic space to carry the instance `is_doubling_measure volume` but volumes grow -exponentially in hyperbolic space. To be really explicit, consider the hyperbolic plane of -curvature -1, the area of a disc of radius `ε` is `A(ε) = 2π(cosh(ε) - 1)` so `A(2ε)/A(ε) ~ exp(ε)`. --/ -class is_doubling_measure {α : Type*} [metric_space α] [measurable_space α] (μ : measure α) := +example we want hyperbolic space to carry the instance `is_unif_loc_doubling_measure volume` but +volumes grow exponentially in hyperbolic space. To be really explicit, consider the hyperbolic plane +of curvature -1, the area of a disc of radius `ε` is `A(ε) = 2π(cosh(ε) - 1)` so +`A(2ε)/A(ε) ~ exp(ε)`. -/ +class is_unif_loc_doubling_measure + {α : Type*} [metric_space α] [measurable_space α] (μ : measure α) := (exists_measure_closed_ball_le_mul [] : ∃ (C : ℝ≥0), ∀ᶠ ε in 𝓝[>] 0, ∀ x, μ (closed_ball x (2 * ε)) ≤ C * μ (closed_ball x ε)) -namespace is_doubling_measure +namespace is_unif_loc_doubling_measure -variables {α : Type*} [metric_space α] [measurable_space α] (μ : measure α) [is_doubling_measure μ] +variables {α : Type*} [metric_space α] [measurable_space α] + (μ : measure α) [is_unif_loc_doubling_measure μ] -/-- A doubling constant for a doubling measure. +/-- A doubling constant for a uniformly locally doubling measure. -See also `is_doubling_measure.scaling_constant_of`. -/ +See also `is_unif_loc_doubling_measure.scaling_constant_of`. -/ def doubling_constant : ℝ≥0 := classical.some $ exists_measure_closed_ball_le_mul μ lemma exists_measure_closed_ball_le_mul' : @@ -82,8 +85,8 @@ begin exact real.rpow_le_rpow_of_exponent_le one_le_two (nat.le_ceil (real.logb 2 K)), }, end -/-- A variant of `is_doubling_measure.doubling_constant` which allows for scaling the radius by -values other than `2`. -/ +/-- A variant of `is_unif_loc_doubling_measure.doubling_constant` which allows for scaling the +radius by values other than `2`. -/ def scaling_constant_of (K : ℝ) : ℝ≥0 := max (classical.some $ exists_eventually_forall_measure_closed_ball_le_mul μ K) 1 @@ -139,4 +142,4 @@ lemma measure_mul_le_scaling_constant_of_mul {K : ℝ} {x : α} {t r : ℝ} μ (closed_ball x (t * r)) ≤ scaling_constant_of μ K * μ (closed_ball x r) := (eventually_measure_mul_le_scaling_constant_of_mul μ K).some_spec.2 x t r ht hr -end is_doubling_measure +end is_unif_loc_doubling_measure diff --git a/src/measure_theory/measure/haar_lebesgue.lean b/src/measure_theory/measure/haar_lebesgue.lean index ebfbc9b84360a..ac4e7f454a807 100644 --- a/src/measure_theory/measure/haar_lebesgue.lean +++ b/src/measure_theory/measure/haar_lebesgue.lean @@ -566,7 +566,8 @@ calc { simp only [ennreal.of_real_ne_top, ne.def, not_false_iff] } end -@[priority 100] instance is_doubling_measure_of_is_add_haar_measure : is_doubling_measure μ := +@[priority 100] instance is_unif_loc_doubling_measure_of_is_add_haar_measure : + is_unif_loc_doubling_measure μ := begin refine ⟨⟨(2 : ℝ≥0) ^ (finrank ℝ E), _⟩⟩, filter_upwards [self_mem_nhds_within] with r hr x, From 500ccb102e657148301f2ec059e8ee766b8660f3 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sun, 2 Apr 2023 16:48:25 +0000 Subject: [PATCH 0756/1291] chore(linear_algebra/dimension): remove a nontriviality assumption (#18715) `dim_pos` does not need `nontrivial R`. Also adds a new lemma that demonstrates why the assumption is still needed on some later results. --- src/linear_algebra/dimension.lean | 45 ++++++++++++++++++++++++------- 1 file changed, 36 insertions(+), 9 deletions(-) diff --git a/src/linear_algebra/dimension.lean b/src/linear_algebra/dimension.lean index 6ccf5e71b30f0..1dc84fa2970ef 100644 --- a/src/linear_algebra/dimension.lean +++ b/src/linear_algebra/dimension.lean @@ -463,23 +463,53 @@ end section rank_zero variables {R : Type u} {M : Type v} -variables [ring R] [nontrivial R] [add_comm_group M] [module R M] [no_zero_smul_divisors R M] +variables [ring R] [add_comm_group M] [module R M] + +@[simp] lemma dim_subsingleton [subsingleton R] : module.rank R M = 1 := +begin + haveI := module.subsingleton R M, + haveI : nonempty {s : set M // linear_independent R (coe : s → M)}, + { exact ⟨⟨∅, linear_independent_empty _ _⟩⟩ }, + rw [module.rank, csupr_eq_of_forall_le_of_forall_lt_exists_gt], + { rintros ⟨s, hs⟩, + rw cardinal.mk_le_one_iff_set_subsingleton, + apply subsingleton_of_subsingleton }, + intros w hw, + refine ⟨⟨{0}, _⟩, _⟩, + { rw linear_independent_iff', + intros, + exact subsingleton.elim _ _ }, + { exact hw.trans_eq (cardinal.mk_singleton _).symm }, +end + +variables [no_zero_smul_divisors R M] + +lemma dim_pos [nontrivial M] : 0 < module.rank R M := +begin + obtain ⟨x, hx⟩ := exists_ne (0 : M), + suffices : 1 ≤ module.rank R M, + { exact zero_lt_one.trans_le this }, + letI := module.nontrivial R M, + suffices : linear_independent R (λ (y : ({x} : set M)), ↑y), + { simpa using (cardinal_le_dim_of_linear_independent this), }, + exact linear_independent_singleton hx +end + +variables [nontrivial R] lemma dim_zero_iff_forall_zero : module.rank R M = 0 ↔ ∀ x : M, x = 0 := begin refine ⟨λ h, _, λ h, _⟩, { contrapose! h, obtain ⟨x, hx⟩ := h, - suffices : 1 ≤ module.rank R M, - { intro h, exact this.not_lt (h.symm ▸ zero_lt_one) }, - suffices : linear_independent R (λ (y : ({x} : set M)), ↑y), - { simpa using (cardinal_le_dim_of_linear_independent this), }, - exact linear_independent_singleton hx }, + letI : nontrivial M := nontrivial_of_ne _ _ hx, + exact dim_pos.ne' }, { have : (⊤ : submodule R M) = ⊥, { ext x, simp [h x] }, rw [←dim_top, this, dim_bot] } end +/-- See `dim_subsingleton` for the reason that `nontrivial R` is needed. -/ lemma dim_zero_iff : module.rank R M = 0 ↔ subsingleton M := dim_zero_iff_forall_zero.trans (subsingleton_iff_forall_eq 0).symm @@ -492,9 +522,6 @@ end lemma dim_pos_iff_nontrivial : 0 < module.rank R M ↔ nontrivial M := dim_pos_iff_exists_ne_zero.trans (nontrivial_iff_exists_ne 0).symm -lemma dim_pos [h : nontrivial M] : 0 < module.rank R M := -dim_pos_iff_nontrivial.2 h - end rank_zero section invariant_basis_number From e5820f6c8fcf1b75bcd7738ae4da1c5896191f72 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sun, 2 Apr 2023 18:44:55 +0000 Subject: [PATCH 0757/1291] chore(algebra/ring_quot): link ring_quot.rel with ring_con_gen (#17892) It's not clear to me whether it's worth keeping `ring_quot.rel r` around, or if it would be better to replace it entirely with `ring_con_gen`. --- src/algebra/ring_quot.lean | 74 ++++++++++++++++++++++++++++++++++++++ 1 file changed, 74 insertions(+) diff --git a/src/algebra/ring_quot.lean b/src/algebra/ring_quot.lean index 231de20a65f7a..e153968e3fcfb 100644 --- a/src/algebra/ring_quot.lean +++ b/src/algebra/ring_quot.lean @@ -27,6 +27,19 @@ variables {R : Type u₁} [semiring R] variables {S : Type u₂} [comm_semiring S] variables {A : Type u₃} [semiring A] [algebra S A] +namespace ring_con + +instance (c : ring_con A) : algebra S c.quotient := +{ smul := (•), + to_ring_hom := c.mk'.comp (algebra_map S A), + commutes' := λ r, quotient.ind' $ by exact λ a, congr_arg quotient.mk' $ algebra.commutes _ _, + smul_def' := λ r, quotient.ind' $ by exact λ a, congr_arg quotient.mk' $ algebra.smul_def _ _ } + +@[simp, norm_cast] lemma coe_algebra_map (c : ring_con A) (s : S) : + (↑(algebra_map S A s) : c.quotient) = algebra_map S _ s := rfl + +end ring_con + namespace ring_quot /-- @@ -58,6 +71,67 @@ by simp only [sub_eq_add_neg, h.neg.add_right] theorem rel.smul {r : A → A → Prop} (k : S) ⦃a b : A⦄ (h : rel r a b) : rel r (k • a) (k • b) := by simp only [algebra.smul_def, rel.mul_right h] +/-- `eqv_gen (ring_quot.rel r)` is a ring congruence. -/ +def ring_con (r : R → R → Prop) : ring_con R := +{ r := eqv_gen (rel r), + iseqv := eqv_gen.is_equivalence _, + add' := λ a b c d hab hcd, begin + induction hab with a' b' hab e a' b' hab' _ c' d' e hcd' hde' ihcd' ihde' generalizing c d, + { refine (eqv_gen.rel _ _ hab.add_left).trans _ _ _ _, + induction hcd with c' d' hcd f c' d' hcd' habcd' c' d' f' hcd' hdf' hbcd' hbcf', + { exact (eqv_gen.rel _ _ hcd.add_right), }, + { exact (eqv_gen.refl _), }, + { exact (habcd'.symm _ _), }, + { exact hbcd'.trans _ _ _ hbcf', }, }, + { induction hcd with c' d' hcd f c' d' hcd' habcd' c' d' f' hcd' hdf' hbcd' hbcf', + { exact (eqv_gen.rel _ _ hcd.add_right), }, + { exact (eqv_gen.refl _), }, + { exact (eqv_gen.symm _ _ habcd'), }, + { exact hbcd'.trans _ _ _ hbcf' }, }, + { exact (hab_ih _ _ $ hcd.symm _ _).symm _ _, }, + { exact (ihcd' _ _ hcd).trans _ _ _ (ihde' _ _ $ eqv_gen.refl _), }, + end, + mul' := λ a b c d hab hcd, begin + induction hab with a' b' hab e a' b' hab' _ c' d' e hcd' hde' ihcd' ihde' generalizing c d, + { refine (eqv_gen.rel _ _ hab.mul_left).trans _ _ _ _, + induction hcd with c' d' hcd f c' d' hcd' habcd' c' d' f' hcd' hdf' hbcd' hbcf', + { exact (eqv_gen.rel _ _ hcd.mul_right), }, + { exact (eqv_gen.refl _), }, + { exact (habcd'.symm _ _), }, + { exact hbcd'.trans _ _ _ hbcf', }, }, + { induction hcd with c' d' hcd f c' d' hcd' habcd' c' d' f' hcd' hdf' hbcd' hbcf', + { exact (eqv_gen.rel _ _ hcd.mul_right), }, + { exact (eqv_gen.refl _), }, + { exact (eqv_gen.symm _ _ habcd'), }, + { exact hbcd'.trans _ _ _ hbcf' }, }, + { exact (hab_ih _ _ $ hcd.symm _ _).symm _ _, }, + { exact (ihcd' _ _ hcd).trans _ _ _ (ihde' _ _ $ eqv_gen.refl _), }, + end } + +lemma eqv_gen_rel_eq (r : R → R → Prop) : eqv_gen (rel r) = ring_con_gen.rel r := +begin + ext x₁ x₂, + split, + { intro h, + induction h with x₃ x₄ h₃₄, + { induction h₃₄ with _ dfg h₃₄ x₃ x₄ x₅ h₃₄', + { exact ring_con_gen.rel.of _ _ ‹_› }, + { exact h₃₄_ih.add (ring_con_gen.rel.refl _) }, + { exact h₃₄_ih.mul (ring_con_gen.rel.refl _) }, + { exact (ring_con_gen.rel.refl _).mul h₃₄_ih} }, + { exact ring_con_gen.rel.refl _ }, + { exact ring_con_gen.rel.symm ‹_› }, + { exact ring_con_gen.rel.trans ‹_› ‹_› } }, + { intro h, + induction h, + { exact eqv_gen.rel _ _ (rel.of ‹_›), }, + { exact (ring_quot.ring_con r).refl _, }, + { exact (ring_quot.ring_con r).symm ‹_›, }, + { exact (ring_quot.ring_con r).trans ‹_› ‹_›, }, + { exact (ring_quot.ring_con r).add ‹_› ‹_›, }, + { exact (ring_quot.ring_con r).mul ‹_› ‹_›, } } +end + end ring_quot /-- The quotient of a ring by an arbitrary relation. -/ From d95bef0d215ea58c0fd7bbc4b151bf3fe952c095 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sun, 2 Apr 2023 18:44:57 +0000 Subject: [PATCH 0758/1291] feat(data/matrix/reflection): lemmas for arbitrary concrete matrices, proved via reflection (#18711) Split from #15738. This contains no meta code, so should be straightforward to port to mathlib4. --- src/data/fin/tuple/reflection.lean | 141 ++++++++++++++++++++++ src/data/matrix/reflection.lean | 185 +++++++++++++++++++++++++++++ test/matrix_reflection.lean | 19 +++ 3 files changed, 345 insertions(+) create mode 100644 src/data/fin/tuple/reflection.lean create mode 100644 src/data/matrix/reflection.lean create mode 100644 test/matrix_reflection.lean diff --git a/src/data/fin/tuple/reflection.lean b/src/data/fin/tuple/reflection.lean new file mode 100644 index 0000000000000..413c3ef0f4ebe --- /dev/null +++ b/src/data/fin/tuple/reflection.lean @@ -0,0 +1,141 @@ +/- +Copyright (c) 2022 Eric Wieser. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Eric Wieser +-/ +import data.fin.vec_notation +import algebra.big_operators.fin + +/-! +# Lemmas for tuples `fin m → α` + +This file contains alternative definitions of common operators on vectors which expand +definitionally to the expected expression when evaluated on `![]` notation. + +This allows "proof by reflection", where we prove `f = ![f 0, f 1]` by defining +`fin_vec.eta_expand f` to be equal to the RHS definitionally, and then prove that +`f = eta_expand f`. + +The definitions in this file should normally not be used directly; the intent is for the +corresponding `*_eq` lemmas to be used in a place where they are definitionally unfolded. + +## Main definitions + +* `fin_vec.seq` +* `fin_vec.map` +* `fin_vec.sum` +* `fin_vec.eta_expand` +-/ + +namespace fin_vec +variables {m n : ℕ} {α β γ : Type*} + +/-- Evaluate `fin_vec.seq f v = ![(f 0) (v 0), (f 1) (v 1), ...]` -/ +def seq : Π {m}, (fin m → (α → β)) → (fin m → α) → fin m → β +| 0 f v := ![] +| (n + 1) f v := matrix.vec_cons (f 0 (v 0)) (seq (matrix.vec_tail f) (matrix.vec_tail v)) + +@[simp] +lemma seq_eq : Π {m} (f : fin m → (α → β)) (v : fin m → α), + seq f v = (λ i, f i (v i)) +| 0 f v := subsingleton.elim _ _ +| (n + 1) f v := funext $ λ i, begin + simp_rw [seq, seq_eq], + refine i.cases _ (λ i, _), + { refl }, + { simp only [matrix.cons_val_succ], refl }, +end + +example {f₁ f₂ : α → β} (a₁ a₂ : α) : seq ![f₁, f₂] ![a₁, a₂] = ![f₁ a₁, f₂ a₂] := rfl + +/-- `fin_vec.map f v = ![f (v 0), f (v 1), ...]` -/ +def map (f : α → β) {m} : (fin m → α) → fin m → β := seq (λ i, f) + +/-- +This can be use to prove +```lean +example {f : α → β} (a₁ a₂ : α) : f ∘ ![a₁, a₂] = ![f a₁, f a₂] := +(map_eq _ _).symm +``` +-/ +@[simp] +lemma map_eq (f : α → β) {m} (v : fin m → α) : map f v = (f ∘ v) := +seq_eq _ _ + +example {f : α → β} (a₁ a₂ : α) : f ∘ ![a₁, a₂] = ![f a₁, f a₂] := +(map_eq _ _).symm + +/-- Expand `v` to `![v 0, v 1, ...]` -/ +def eta_expand {m} (v : fin m → α) : fin m → α := map id v + +/-- +This can be use to prove +```lean +example {f : α → β} (a : fin 2 → α) : a = ![a 0, a 1] := (eta_expand_eq _).symm +``` +-/ +@[simp] +lemma eta_expand_eq {m} (v : fin m → α) : eta_expand v = v := map_eq id v + +example {f : α → β} (a : fin 2 → α) : a = ![a 0, a 1] := (eta_expand_eq _).symm + +/-- `∀` with better defeq for `∀ x : fin m → α, P x`. -/ +def «forall» : Π {m} (P : (fin m → α) → Prop), Prop +| 0 P := P ![] +| (n + 1) P := ∀ x : α, «forall» (λ v, P (matrix.vec_cons x v)) + +/-- +This can be use to prove +```lean +example (P : (fin 2 → α) → Prop) : (∀ f, P f) ↔ (∀ a₀ a₁, P ![a₀, a₁]) := (forall_iff _).symm +``` +-/ +@[simp] +lemma forall_iff : Π {m} (P : (fin m → α) → Prop), «forall» P ↔ ∀ x, P x +| 0 P := by { simp only [«forall», fin.forall_fin_zero_pi], refl } +| (n + 1) P := by simp only [«forall», forall_iff, fin.forall_fin_succ_pi, matrix.vec_cons] + +example (P : (fin 2 → α) → Prop) : (∀ f, P f) ↔ (∀ a₀ a₁, P ![a₀, a₁]) := (forall_iff _).symm + +/-- `∃` with better defeq for `∃ x : fin m → α, P x`. -/ +def «exists» : Π {m} (P : (fin m → α) → Prop), Prop +| 0 P := P ![] +| (n + 1) P := ∃ x : α, «exists» (λ v, P (matrix.vec_cons x v)) + +/-- +This can be use to prove +```lean +example (P : (fin 2 → α) → Prop) : (∃ f, P f) ↔ (∃ a₀ a₁, P ![a₀, a₁]) := (exists_iff _).symm +``` +-/ +lemma exists_iff : Π {m} (P : (fin m → α) → Prop), «exists» P ↔ ∃ x, P x +| 0 P := by { simp only [«exists», fin.exists_fin_zero_pi, matrix.vec_empty], refl } +| (n + 1) P := by simp only [«exists», exists_iff, fin.exists_fin_succ_pi, matrix.vec_cons] + +example (P : (fin 2 → α) → Prop) : (∃ f, P f) ↔ (∃ a₀ a₁, P ![a₀, a₁]) := (exists_iff _).symm + +/-- `finset.univ.sum` with better defeq for `fin` -/ +def sum [has_add α] [has_zero α] : Π{m} (v : fin m → α), α +| 0 v := 0 +| 1 v := v 0 +| (n + 2) v := sum (v ∘ fin.cast_succ) + v (fin.last _) + +open_locale big_operators + +/-- This can be used to prove +```lean +example [add_comm_monoid α] (a : fin 3 → α) : ∑ i, a i = a 0 + a 1 + a 2 := +(sum_eq _).symm +``` +-/ +@[simp] +lemma sum_eq [add_comm_monoid α] : Π {m} (a : fin m → α), + sum a = ∑ i, a i +| 0 a := rfl +| 1 a := (fintype.sum_unique a).symm +| (n + 2) a := by rw [fin.sum_univ_cast_succ, sum, sum_eq] + +example [add_comm_monoid α] (a : fin 3 → α) : ∑ i, a i = a 0 + a 1 + a 2 := +(sum_eq _).symm + +end fin_vec diff --git a/src/data/matrix/reflection.lean b/src/data/matrix/reflection.lean new file mode 100644 index 0000000000000..ea0f5921db200 --- /dev/null +++ b/src/data/matrix/reflection.lean @@ -0,0 +1,185 @@ +/- +Copyright (c) 2022 Eric Wieser. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Eric Wieser +-/ +import data.matrix.notation +import data.matrix.basic +import data.fin.tuple.reflection + +/-! +# Lemmas for concrete matrices `matrix (fin m) (fin n) α` + +This file contains alternative definitions of common operators on matrices that expand +definitionally to the expected expression when evaluated on `!![]` notation. + +This allows "proof by reflection", where we prove `A = !![A 0 0, A 0 1; A 1 0, A 1 1]` by defining +`matrix.eta_expand A` to be equal to the RHS definitionally, and then prove that +`A = eta_expand A`. + +The definitions in this file should normally not be used directly; the intent is for the +corresponding `*_eq` lemmas to be used in a place where they are definitionally unfolded. + +## Main definitionss + +* `matrix.transposeᵣ` +* `matrix.dot_productᵣ` +* `matrix.mulᵣ` +* `matrix.eta_expand` + +-/ + +open_locale matrix + +namespace matrix +variables {l m n : ℕ} {α β : Type*} + +/-- `∀` with better defeq for `∀ x : matrix (fin m) (fin n) α, P x`. -/ +def «forall» : Π {m n} (P : (matrix (fin m) (fin n) α) → Prop), Prop +| 0 n P := P (of ![]) +| (m + 1) n P := fin_vec.forall $ λ r, «forall» (λ A, P (of (matrix.vec_cons r A))) + +/-- +This can be use to prove +```lean +example (P : matrix (fin 2) (fin 3) α → Prop) : + (∀ x, P x) ↔ ∀ a b c d e f, P !![a, b, c; d, e, f] := +(forall_iff _).symm +``` +-/ +lemma forall_iff : Π {m n} (P : (matrix (fin m) (fin n) α) → Prop), «forall» P ↔ ∀ x, P x +| 0 n P := iff.symm fin.forall_fin_zero_pi +| (m + 1) n P := begin + simp only [«forall», fin_vec.forall_iff, forall_iff], + exact iff.symm fin.forall_fin_succ_pi, +end + +example (P : matrix (fin 2) (fin 3) α → Prop) : + (∀ x, P x) ↔ ∀ a b c d e f, P !![a, b, c; d, e, f] := +(forall_iff _).symm + +/--`∃` with better defeq for `∃ x : matrix (fin m) (fin n) α, P x`. -/ +def «exists» : Π {m n} (P : (matrix (fin m) (fin n) α) → Prop), Prop +| 0 n P := P (of ![]) +| (m + 1) n P := fin_vec.exists $ λ r, «exists» (λ A, P (of (matrix.vec_cons r A))) + +/-- +This can be use to prove +```lean +example (P : matrix (fin 2) (fin 3) α → Prop) : + (∃ x, P x) ↔ ∃ a b c d e f, P !![a, b, c; d, e, f] := +(exists_iff _).symm +``` +-/ +lemma exists_iff : Π {m n} (P : (matrix (fin m) (fin n) α) → Prop), «exists» P ↔ ∃ x, P x +| 0 n P := iff.symm fin.exists_fin_zero_pi +| (m + 1) n P := begin + simp only [«exists», fin_vec.exists_iff, exists_iff], + exact iff.symm fin.exists_fin_succ_pi, +end + +example (P : matrix (fin 2) (fin 3) α → Prop) : + (∃ x, P x) ↔ ∃ a b c d e f, P !![a, b, c; d, e, f] := +(exists_iff _).symm + +/-- `matrix.tranpose` with better defeq for `fin` -/ +def transposeᵣ : Π {m n}, matrix (fin m) (fin n) α → matrix (fin n) (fin m) α +| _ 0 A := of ![] +| m (n + 1) A := of $ vec_cons (fin_vec.map (λ v : fin _ → α, v 0) A) + (transposeᵣ (A.submatrix id fin.succ)) + +/-- This can be used to prove +```lean +example (a b c d : α) : transpose !![a, b; c, d] = !![a, c; b, d] := (transposeᵣ_eq _).symm +``` +-/ +@[simp] +lemma transposeᵣ_eq : Π {m n} (A : matrix (fin m) (fin n) α), + transposeᵣ A = transpose A +| _ 0 A := subsingleton.elim _ _ +| m (n + 1) A := matrix.ext $ λ i j, begin + simp_rw [transposeᵣ, transposeᵣ_eq], + refine i.cases _ (λ i, _), + { dsimp, rw fin_vec.map_eq }, + { simp only [of_apply, matrix.cons_val_succ], refl }, +end + +example (a b c d : α) : transpose !![a, b; c, d] = !![a, c; b, d] := (transposeᵣ_eq _).symm + +/-- `matrix.dot_product` with better defeq for `fin` -/ +def dot_productᵣ [has_mul α] [has_add α] [has_zero α] {m} (a b : fin m → α) : α := +fin_vec.sum $ fin_vec.seq (fin_vec.map (*) a) b + +/-- This can be used to prove +```lean +example (a b c d : α) [has_mul α] [add_comm_monoid α] : + dot_product ![a, b] ![c, d] = a * c + b * d := +(dot_productᵣ_eq _ _).symm +``` +-/ +@[simp] +lemma dot_productᵣ_eq [has_mul α] [add_comm_monoid α] {m} (a b : fin m → α) : + dot_productᵣ a b = dot_product a b := +by simp_rw [dot_productᵣ, dot_product, fin_vec.sum_eq, fin_vec.seq_eq, fin_vec.map_eq] + +example (a b c d : α) [has_mul α] [add_comm_monoid α] : + dot_product ![a, b] ![c, d] = a * c + b * d := +(dot_productᵣ_eq _ _).symm + +/-- `matrix.mul` with better defeq for `fin` -/ +def mulᵣ [has_mul α] [has_add α] [has_zero α] + (A : matrix (fin l) (fin m) α) (B : matrix (fin m) (fin n) α) : + matrix (fin l) (fin n) α := +of $ fin_vec.map (λ v₁, fin_vec.map (λ v₂, dot_productᵣ v₁ v₂) (transposeᵣ B)) A + +/-- This can be used to prove +```lean +example [add_comm_monoid α] [has_mul α] + (a₁₁ a₁₂ a₁₃ a₂₁ a₂₂ a₂₃ a₃₁ a₃₂ a₃₃ b₁₁ b₁₂ b₁₃ b₂₁ b₂₂ b₂₃ b₃₁ b₃₂ b₃₃ : α) : + !![a₁₁, a₁₂; + a₂₁, a₂₂] ⬝ !![b₁₁, b₁₂; + b₂₁, b₂₂] = + !![a₁₁*b₁₁ + a₁₂*b₂₁, a₁₁*b₁₂ + a₁₂*b₂₂; + a₂₁*b₁₁ + a₂₂*b₂₁, a₂₁*b₁₂ + a₂₂*b₂₂] := +``` +-/ +@[simp] +lemma mulᵣ_eq [has_mul α] [add_comm_monoid α] + (A : matrix (fin l) (fin m) α) (B : matrix (fin m) (fin n) α) : + mulᵣ A B = A.mul B := +begin + simp [mulᵣ, function.comp, matrix.mul, matrix.transpose], + refl, +end + +example [add_comm_monoid α] [has_mul α] + (a₁₁ a₁₂ a₁₃ a₂₁ a₂₂ a₂₃ a₃₁ a₃₂ a₃₃ b₁₁ b₁₂ b₁₃ b₂₁ b₂₂ b₂₃ b₃₁ b₃₂ b₃₃ : α) : + !![a₁₁, a₁₂; + a₂₁, a₂₂].mul !![b₁₁, b₁₂; + b₂₁, b₂₂] = + !![a₁₁*b₁₁ + a₁₂*b₂₁, a₁₁*b₁₂ + a₁₂*b₂₂; + a₂₁*b₁₁ + a₂₂*b₂₁, a₂₁*b₁₂ + a₂₂*b₂₂] := +(mulᵣ_eq _ _).symm + +/-- Expand `A` to `!![A 0 0, ...; ..., A m n]` -/ +def eta_expand {m n} (A : matrix (fin m) (fin n) α) : matrix (fin m) (fin n) α := +matrix.of (fin_vec.eta_expand (λ i, fin_vec.eta_expand (λ j, A i j))) + +/-- This can be used to prove +```lean +example (A : matrix (fin 2) (fin 2) α) : + A = !![A 0 0, A 0 1; + A 1 0, A 1 1] := +(eta_expand_eq _).symm +``` +-/ +lemma eta_expand_eq {m n} (A : matrix (fin m) (fin n) α) : + eta_expand A = A := +by simp_rw [eta_expand, fin_vec.eta_expand_eq, matrix.of, equiv.refl_apply] + +example (A : matrix (fin 2) (fin 2) α) : + A = !![A 0 0, A 0 1; + A 1 0, A 1 1] := +(eta_expand_eq _).symm + +end matrix diff --git a/test/matrix_reflection.lean b/test/matrix_reflection.lean new file mode 100644 index 0000000000000..331d0a5573bfc --- /dev/null +++ b/test/matrix_reflection.lean @@ -0,0 +1,19 @@ +import data.matrix.reflection + +variables {α: Type*} + +open_locale matrix +open matrix + +-- This one is too long for the docstring, but is nice to have tested +example [add_comm_monoid α] [has_mul α] + (a₁₁ a₁₂ a₁₃ a₂₁ a₂₂ a₂₃ a₃₁ a₃₂ a₃₃ b₁₁ b₁₂ b₁₃ b₂₁ b₂₂ b₂₃ b₃₁ b₃₂ b₃₃ : α) : + !![a₁₁, a₁₂, a₁₃; + a₂₁, a₂₂, a₂₃; + a₃₁, a₃₂, a₃₃] ⬝ !![b₁₁, b₁₂, b₁₃; + b₂₁, b₂₂, b₂₃; + b₃₁, b₃₂, b₃₃] = + !![a₁₁*b₁₁ + a₁₂*b₂₁ + a₁₃*b₃₁, a₁₁*b₁₂ + a₁₂*b₂₂ + a₁₃*b₃₂, a₁₁*b₁₃ + a₁₂*b₂₃ + a₁₃*b₃₃; + a₂₁*b₁₁ + a₂₂*b₂₁ + a₂₃*b₃₁, a₂₁*b₁₂ + a₂₂*b₂₂ + a₂₃*b₃₂, a₂₁*b₁₃ + a₂₂*b₂₃ + a₂₃*b₃₃; + a₃₁*b₁₁ + a₃₂*b₂₁ + a₃₃*b₃₁, a₃₁*b₁₂ + a₃₂*b₂₂ + a₃₃*b₃₂, a₃₁*b₁₃ + a₃₂*b₂₃ + a₃₃*b₃₃] := +(matrix.mulᵣ_eq _ _).symm From 31ca6f9cf5f90a6206092cd7f84b359dcb6d52e0 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Sun, 2 Apr 2023 19:51:20 +0000 Subject: [PATCH 0759/1291] chore(*): add mathlib4 synchronization comments (#18662) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.category.Group.preadditive` * `algebra.category.Group.zero` * `algebra.char_p.quotient` * `algebra.cubic_discriminant` * `algebra.gcd_monoid.div` * `algebra.monoid_algebra.division` * `algebra.star.chsh` * `analysis.convex.basic` * `analysis.convex.extreme` * `analysis.convex.hull` * `analysis.convex.strict` * `category_theory.adjunction.opposites` * `category_theory.category.Twop` * `category_theory.elements` * `category_theory.glue_data` * `category_theory.localization.opposite` * `category_theory.localization.predicate` * `category_theory.monoidal.discrete` * `category_theory.monoidal.linear` * `category_theory.monoidal.transport` * `category_theory.sites.pretopology` * `category_theory.sites.spaces` * `combinatorics.derangements.finite` * `combinatorics.simple_graph.matching` * `computability.primrec` * `data.mllist` * `data.mv_polynomial.cardinal` * `data.mv_polynomial.division` * `data.nat.choose.factorization` * `data.nat.factorization.basic` * `data.nat.factorization.prime_pow` * `data.nat.totient` * `data.polynomial.degree.card_pow_degree` * `data.polynomial.mirror` * `data.polynomial.partial_fractions` * `data.polynomial.splits` * `data.set.list` * `data.set.pairwise.lattice` * `linear_algebra.affine_space.midpoint_zero` * `linear_algebra.affine_space.ordered` * `linear_algebra.affine_space.pointwise` * `linear_algebra.direct_sum.finsupp` * `linear_algebra.direct_sum.tensor_product` * `model_theory.basic` * `number_theory.lucas_primality` * `order.upper_lower.locally_finite` * `representation_theory.maschke` * `ring_theory.adjoin.fg` * `ring_theory.localization.module` * `ring_theory.mv_polynomial.weighted_homogeneous` * `ring_theory.polynomial.basic` * `ring_theory.polynomial.vieta` * `topology.algebra.polynomial` --- src/algebra/category/Group/preadditive.lean | 3 +++ src/algebra/category/Group/zero.lean | 3 +++ src/algebra/char_p/quotient.lean | 3 +++ src/algebra/cubic_discriminant.lean | 3 +++ src/algebra/gcd_monoid/div.lean | 3 +++ src/algebra/monoid_algebra/division.lean | 3 +++ src/algebra/star/chsh.lean | 3 +++ src/analysis/convex/basic.lean | 3 +++ src/analysis/convex/extreme.lean | 3 +++ src/analysis/convex/hull.lean | 3 +++ src/analysis/convex/strict.lean | 3 +++ src/category_theory/adjunction/opposites.lean | 3 +++ src/category_theory/category/Twop.lean | 3 +++ src/category_theory/elements.lean | 3 +++ src/category_theory/glue_data.lean | 3 +++ src/category_theory/localization/opposite.lean | 3 +++ src/category_theory/localization/predicate.lean | 3 +++ src/category_theory/monoidal/discrete.lean | 3 +++ src/category_theory/monoidal/linear.lean | 3 +++ src/category_theory/monoidal/transport.lean | 3 +++ src/category_theory/sites/pretopology.lean | 3 +++ src/category_theory/sites/spaces.lean | 3 +++ src/combinatorics/derangements/finite.lean | 3 +++ src/combinatorics/simple_graph/matching.lean | 3 +++ src/computability/primrec.lean | 3 +++ src/data/mllist.lean | 3 +++ src/data/mv_polynomial/cardinal.lean | 3 +++ src/data/mv_polynomial/division.lean | 3 +++ src/data/nat/choose/factorization.lean | 3 +++ src/data/nat/factorization/basic.lean | 3 +++ src/data/nat/factorization/prime_pow.lean | 3 +++ src/data/nat/totient.lean | 3 +++ src/data/polynomial/degree/card_pow_degree.lean | 3 +++ src/data/polynomial/mirror.lean | 3 +++ src/data/polynomial/partial_fractions.lean | 3 +++ src/data/polynomial/splits.lean | 3 +++ src/data/set/list.lean | 3 +++ src/data/set/pairwise/lattice.lean | 3 +++ src/linear_algebra/affine_space/midpoint_zero.lean | 3 +++ src/linear_algebra/affine_space/ordered.lean | 3 +++ src/linear_algebra/affine_space/pointwise.lean | 3 +++ src/linear_algebra/direct_sum/finsupp.lean | 3 +++ src/linear_algebra/direct_sum/tensor_product.lean | 3 +++ src/model_theory/basic.lean | 3 +++ src/number_theory/lucas_primality.lean | 3 +++ src/order/upper_lower/locally_finite.lean | 3 +++ src/representation_theory/maschke.lean | 3 +++ src/ring_theory/adjoin/fg.lean | 3 +++ src/ring_theory/localization/module.lean | 3 +++ src/ring_theory/mv_polynomial/weighted_homogeneous.lean | 3 +++ src/ring_theory/polynomial/basic.lean | 3 +++ src/ring_theory/polynomial/vieta.lean | 3 +++ src/topology/algebra/polynomial.lean | 3 +++ 53 files changed, 159 insertions(+) diff --git a/src/algebra/category/Group/preadditive.lean b/src/algebra/category/Group/preadditive.lean index 04317799f588e..f6002036a76f0 100644 --- a/src/algebra/category/Group/preadditive.lean +++ b/src/algebra/category/Group/preadditive.lean @@ -8,6 +8,9 @@ import category_theory.preadditive.basic /-! # The category of additive commutative groups is preadditive. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ open category_theory diff --git a/src/algebra/category/Group/zero.lean b/src/algebra/category/Group/zero.lean index f16544873c4d5..a58fbfeb33a4b 100644 --- a/src/algebra/category/Group/zero.lean +++ b/src/algebra/category/Group/zero.lean @@ -9,6 +9,9 @@ import category_theory.limits.shapes.zero_objects /-! # The category of (commutative) (additive) groups has a zero object. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + `AddCommGroup` also has zero morphisms. For definitional reasons, we infer this from preadditivity rather than from the existence of a zero object. -/ diff --git a/src/algebra/char_p/quotient.lean b/src/algebra/char_p/quotient.lean index 922b91b1b5740..d6257701ef9c1 100644 --- a/src/algebra/char_p/quotient.lean +++ b/src/algebra/char_p/quotient.lean @@ -9,6 +9,9 @@ import ring_theory.ideal.quotient /-! # Characteristic of quotients rings + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ universes u v diff --git a/src/algebra/cubic_discriminant.lean b/src/algebra/cubic_discriminant.lean index e10263ae6e55c..505ffb3079eb1 100644 --- a/src/algebra/cubic_discriminant.lean +++ b/src/algebra/cubic_discriminant.lean @@ -9,6 +9,9 @@ import data.polynomial.splits /-! # Cubics and discriminants +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines cubic polynomials over a semiring and their discriminants over a splitting field. ## Main definitions diff --git a/src/algebra/gcd_monoid/div.lean b/src/algebra/gcd_monoid/div.lean index 2c7616974c314..bddd821ee1563 100644 --- a/src/algebra/gcd_monoid/div.lean +++ b/src/algebra/gcd_monoid/div.lean @@ -11,6 +11,9 @@ import ring_theory.polynomial.content /-! # Basic results about setwise gcds on normalized gcd monoid with a division. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main results * `finset.nat.gcd_div_eq_one`: given a nonempty finset `s` and a function `f` from `s` to diff --git a/src/algebra/monoid_algebra/division.lean b/src/algebra/monoid_algebra/division.lean index 3e34e02a80b31..4478367e461e0 100644 --- a/src/algebra/monoid_algebra/division.lean +++ b/src/algebra/monoid_algebra/division.lean @@ -9,6 +9,9 @@ import data.finsupp.order /-! # Division of `add_monoid_algebra` by monomials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file is most important for when `G = ℕ` (polynomials) or `G = σ →₀ ℕ` (multivariate polynomials). diff --git a/src/algebra/star/chsh.lean b/src/algebra/star/chsh.lean index 56a574062a67e..f794817618872 100644 --- a/src/algebra/star/chsh.lean +++ b/src/algebra/star/chsh.lean @@ -9,6 +9,9 @@ import data.real.sqrt /-! # The Clauser-Horne-Shimony-Holt inequality and Tsirelson's inequality. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We establish a version of the Clauser-Horne-Shimony-Holt (CHSH) inequality (which is a generalization of Bell's inequality). This is a foundational result which implies that diff --git a/src/analysis/convex/basic.lean b/src/analysis/convex/basic.lean index 601c2623ecfe3..6825f10078981 100644 --- a/src/analysis/convex/basic.lean +++ b/src/analysis/convex/basic.lean @@ -10,6 +10,9 @@ import linear_algebra.affine_space.affine_subspace /-! # Convex sets and functions in vector spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In a 𝕜-vector space, we define the following objects and properties. * `convex 𝕜 s`: A set `s` is convex if for any two points `x y ∈ s` it includes `segment 𝕜 x y`. * `std_simplex 𝕜 ι`: The standard simplex in `ι → 𝕜` (currently requires `fintype ι`). It is the diff --git a/src/analysis/convex/extreme.lean b/src/analysis/convex/extreme.lean index 772aa6caf5056..810e6ceec6282 100644 --- a/src/analysis/convex/extreme.lean +++ b/src/analysis/convex/extreme.lean @@ -8,6 +8,9 @@ import analysis.convex.hull /-! # Extreme sets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines extreme sets and extreme points for sets in a module. An extreme set of `A` is a subset of `A` that is as far as it can get in any outward direction: If diff --git a/src/analysis/convex/hull.lean b/src/analysis/convex/hull.lean index e31a5497a8504..efbc727029664 100644 --- a/src/analysis/convex/hull.lean +++ b/src/analysis/convex/hull.lean @@ -9,6 +9,9 @@ import order.closure /-! # Convex hull +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the convex hull of a set `s` in a module. `convex_hull 𝕜 s` is the smallest convex set containing `s`. In order theory speak, this is a closure operator. diff --git a/src/analysis/convex/strict.lean b/src/analysis/convex/strict.lean index df09dbab49398..8b9d0655d2936 100644 --- a/src/analysis/convex/strict.lean +++ b/src/analysis/convex/strict.lean @@ -9,6 +9,9 @@ import topology.algebra.order.group /-! # Strictly convex sets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines strictly convex sets. A set is strictly convex if the open segment between any two distinct points lies in its interior. diff --git a/src/category_theory/adjunction/opposites.lean b/src/category_theory/adjunction/opposites.lean index 8566a4b81d802..8a9e9573d4f95 100644 --- a/src/category_theory/adjunction/opposites.lean +++ b/src/category_theory/adjunction/opposites.lean @@ -11,6 +11,9 @@ import category_theory.opposites /-! # Opposite adjunctions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains constructions to relate adjunctions of functors to adjunctions of their opposites. These constructions are used to show uniqueness of adjoints (up to natural isomorphism). diff --git a/src/category_theory/category/Twop.lean b/src/category_theory/category/Twop.lean index af4b02063d683..e5563cc371214 100644 --- a/src/category_theory/category/Twop.lean +++ b/src/category_theory/category/Twop.lean @@ -9,6 +9,9 @@ import data.two_pointing /-! # The category of two-pointed types +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This defines `Twop`, the category of two-pointed types. ## References diff --git a/src/category_theory/elements.lean b/src/category_theory/elements.lean index cfe71eb4176be..5ffe55cca2dc4 100644 --- a/src/category_theory/elements.lean +++ b/src/category_theory/elements.lean @@ -10,6 +10,9 @@ import category_theory.punit /-! # The category of elements +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the category of elements, also known as (a special case of) the Grothendieck construction. diff --git a/src/category_theory/glue_data.lean b/src/category_theory/glue_data.lean index 15b95f35be6b7..94b388c7c0448 100644 --- a/src/category_theory/glue_data.lean +++ b/src/category_theory/glue_data.lean @@ -12,6 +12,9 @@ import category_theory.limits.shapes.types /-! # Gluing data +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define `glue_data` as a family of data needed to glue topological spaces, schemes, etc. We provide the API to realize it as a multispan diagram, and also states lemmas about its interaction with a functor that preserves certain pullbacks. diff --git a/src/category_theory/localization/opposite.lean b/src/category_theory/localization/opposite.lean index f7efcb00d22d4..97b70abd9b354 100644 --- a/src/category_theory/localization/opposite.lean +++ b/src/category_theory/localization/opposite.lean @@ -10,6 +10,9 @@ import category_theory.localization.predicate # Localization of the opposite category +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + If a functor `L : C ⥤ D` is a localization functor for `W : morphism_property C`, it is shown in this file that `L.op : Cᵒᵖ ⥤ Dᵒᵖ` is also a localization functor. diff --git a/src/category_theory/localization/predicate.lean b/src/category_theory/localization/predicate.lean index 3e6a05039647c..89756a2b70089 100644 --- a/src/category_theory/localization/predicate.lean +++ b/src/category_theory/localization/predicate.lean @@ -10,6 +10,9 @@ import category_theory.localization.construction # Predicate for localized categories +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, a predicate `L.is_localization W` is introduced for a functor `L : C ⥤ D` and `W : morphism_property C`: it expresses that `L` identifies `D` with the localized category of `C` with respect to `W` (up to equivalence). diff --git a/src/category_theory/monoidal/discrete.lean b/src/category_theory/monoidal/discrete.lean index f32ed6597d319..45a90c545cf46 100644 --- a/src/category_theory/monoidal/discrete.lean +++ b/src/category_theory/monoidal/discrete.lean @@ -10,6 +10,9 @@ import category_theory.monoidal.natural_transformation /-! # Monoids as discrete monoidal categories +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The discrete category on a monoid is a monoidal category. Multiplicative morphisms induced monoidal functors. -/ diff --git a/src/category_theory/monoidal/linear.lean b/src/category_theory/monoidal/linear.lean index adf3555e1fcec..639762539a871 100644 --- a/src/category_theory/monoidal/linear.lean +++ b/src/category_theory/monoidal/linear.lean @@ -9,6 +9,9 @@ import category_theory.monoidal.preadditive /-! # Linear monoidal categories +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A monoidal category is `monoidal_linear R` if it is monoidal preadditive and tensor product of morphisms is `R`-linear in both factors. -/ diff --git a/src/category_theory/monoidal/transport.lean b/src/category_theory/monoidal/transport.lean index 954619dd50498..d447a53b4ff9c 100644 --- a/src/category_theory/monoidal/transport.lean +++ b/src/category_theory/monoidal/transport.lean @@ -8,6 +8,9 @@ import category_theory.monoidal.natural_transformation /-! # Transport a monoidal structure along an equivalence. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + When `C` and `D` are equivalent as categories, we can transport a monoidal structure on `C` along the equivalence, obtaining a monoidal structure on `D`. diff --git a/src/category_theory/sites/pretopology.lean b/src/category_theory/sites/pretopology.lean index b8d0f9f3fa741..85c835126e84c 100644 --- a/src/category_theory/sites/pretopology.lean +++ b/src/category_theory/sites/pretopology.lean @@ -9,6 +9,9 @@ import category_theory.sites.grothendieck /-! # Grothendieck pretopologies +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Definition and lemmas about Grothendieck pretopologies. A Grothendieck pretopology for a category `C` is a set of families of morphisms with fixed codomain, satisfying certain closure conditions. diff --git a/src/category_theory/sites/spaces.lean b/src/category_theory/sites/spaces.lean index cd20ac5cdb426..8148124c359f5 100644 --- a/src/category_theory/sites/spaces.lean +++ b/src/category_theory/sites/spaces.lean @@ -11,6 +11,9 @@ import topology.sets.opens /-! # Grothendieck topology on a topological space +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Define the Grothendieck topology and the pretopology associated to a topological space, and show that the pretopology induces the topology. diff --git a/src/combinatorics/derangements/finite.lean b/src/combinatorics/derangements/finite.lean index f9aa9e995306a..613acd5a43d7e 100644 --- a/src/combinatorics/derangements/finite.lean +++ b/src/combinatorics/derangements/finite.lean @@ -11,6 +11,9 @@ import tactic.ring /-! # Derangements on fintypes +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains lemmas that describe the cardinality of `derangements α` when `α` is a fintype. # Main definitions diff --git a/src/combinatorics/simple_graph/matching.lean b/src/combinatorics/simple_graph/matching.lean index 7fc17adc471d2..e3803b5e26a3e 100644 --- a/src/combinatorics/simple_graph/matching.lean +++ b/src/combinatorics/simple_graph/matching.lean @@ -9,6 +9,9 @@ import combinatorics.simple_graph.subgraph /-! # Matchings +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A *matching* for a simple graph is a set of disjoint pairs of adjacent vertices, and the set of all the vertices in a matching is called its *support* (and sometimes the vertices in the support are said to be *saturated* by the matching). A *perfect matching* is a matching whose support contains diff --git a/src/computability/primrec.lean b/src/computability/primrec.lean index a6e4b7a34941f..98f60dbe6fa77 100644 --- a/src/computability/primrec.lean +++ b/src/computability/primrec.lean @@ -10,6 +10,9 @@ import logic.function.iterate /-! # The primitive recursive functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The primitive recursive functions are the least collection of functions `nat → nat` which are closed under projections (using the mkpair pairing function), composition, zero, successor, and primitive recursion diff --git a/src/data/mllist.lean b/src/data/mllist.lean index 8fa1b232225cf..be391cfbe0076 100644 --- a/src/data/mllist.lean +++ b/src/data/mllist.lean @@ -7,6 +7,9 @@ import data.option.defs /-! # Monadic lazy lists. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + An alternative construction of lazy lists (see also `data.lazy_list`), with "lazyness" controlled by an arbitrary monad. diff --git a/src/data/mv_polynomial/cardinal.lean b/src/data/mv_polynomial/cardinal.lean index ff7a3335a6bee..cedd2ff85a034 100644 --- a/src/data/mv_polynomial/cardinal.lean +++ b/src/data/mv_polynomial/cardinal.lean @@ -9,6 +9,9 @@ import set_theory.cardinal.ordinal /-! # Cardinality of Multivariate Polynomial Ring +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The main result in this file is `mv_polynomial.cardinal_mk_le_max`, which says that the cardinality of `mv_polynomial σ R` is bounded above by the maximum of `#R`, `#σ` and `ℵ₀`. diff --git a/src/data/mv_polynomial/division.lean b/src/data/mv_polynomial/division.lean index 9e44aca4ab738..5784eda1e59b6 100644 --- a/src/data/mv_polynomial/division.lean +++ b/src/data/mv_polynomial/division.lean @@ -9,6 +9,9 @@ import data.mv_polynomial.basic /-! # Division of `mv_polynomial` by monomials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions * `mv_polynomial.div_monomial x s`: divides `x` by the monomial `mv_polynomial.monomial 1 s` diff --git a/src/data/nat/choose/factorization.lean b/src/data/nat/choose/factorization.lean index 23075edf1b6e7..35268691fd3dd 100644 --- a/src/data/nat/choose/factorization.lean +++ b/src/data/nat/choose/factorization.lean @@ -11,6 +11,9 @@ import data.nat.multiplicity /-! # Factorization of Binomial Coefficients +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains a few results on the multiplicity of prime factors within certain size bounds in binomial coefficients. These include: diff --git a/src/data/nat/factorization/basic.lean b/src/data/nat/factorization/basic.lean index 2e46c11e86c66..492f64d1295d4 100644 --- a/src/data/nat/factorization/basic.lean +++ b/src/data/nat/factorization/basic.lean @@ -13,6 +13,9 @@ import tactic.interval_cases /-! # Prime factorizations +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + `n.factorization` is the finitely supported function `ℕ →₀ ℕ` mapping each prime factor of `n` to its multiplicity in `n`. For example, since 2000 = 2^4 * 5^3, * `factorization 2000 2` is 4 diff --git a/src/data/nat/factorization/prime_pow.lean b/src/data/nat/factorization/prime_pow.lean index b315f2b22beda..ab2df778bbb15 100644 --- a/src/data/nat/factorization/prime_pow.lean +++ b/src/data/nat/factorization/prime_pow.lean @@ -9,6 +9,9 @@ import data.nat.factorization.basic /-! # Prime powers and factorizations +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file deals with factorizations of prime powers. -/ diff --git a/src/data/nat/totient.lean b/src/data/nat/totient.lean index b3082bdf5a895..bd601d3273d33 100644 --- a/src/data/nat/totient.lean +++ b/src/data/nat/totient.lean @@ -11,6 +11,9 @@ import data.zmod.basic /-! # Euler's totient function +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines [Euler's totient function](https://en.wikipedia.org/wiki/Euler's_totient_function) `nat.totient n` which counts the number of naturals less than `n` that are coprime with `n`. We prove the divisor sum formula, namely that `n` equals `φ` summed over the divisors of `n`. See diff --git a/src/data/polynomial/degree/card_pow_degree.lean b/src/data/polynomial/degree/card_pow_degree.lean index d5838fe31283f..0115d6a5f1e2b 100644 --- a/src/data/polynomial/degree/card_pow_degree.lean +++ b/src/data/polynomial/degree/card_pow_degree.lean @@ -9,6 +9,9 @@ import data.polynomial.field_division /-! # Absolute value on polynomials over a finite field. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Let `Fq` be a finite field of cardinality `q`, then the map sending a polynomial `p` to `q ^ degree p` (where `q ^ degree 0 = 0`) is an absolute value. diff --git a/src/data/polynomial/mirror.lean b/src/data/polynomial/mirror.lean index 2c39b8442a113..16b41f1eea6b0 100644 --- a/src/data/polynomial/mirror.lean +++ b/src/data/polynomial/mirror.lean @@ -10,6 +10,9 @@ import data.polynomial.ring_division /-! # "Mirror" of a univariate polynomial +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define `polynomial.mirror`, a variant of `polynomial.reverse`. The difference between `reverse` and `mirror` is that `reverse` will decrease the degree if the polynomial is divisible by `X`. diff --git a/src/data/polynomial/partial_fractions.lean b/src/data/polynomial/partial_fractions.lean index 8192f1d5bc55d..b364b375da840 100644 --- a/src/data/polynomial/partial_fractions.lean +++ b/src/data/polynomial/partial_fractions.lean @@ -14,6 +14,9 @@ import tactic.linear_combination # Partial fractions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + These results were formalised by the Xena Project, at the suggestion of Patrick Massot. diff --git a/src/data/polynomial/splits.lean b/src/data/polynomial/splits.lean index 6d98f5f312386..e833cce8ec92d 100644 --- a/src/data/polynomial/splits.lean +++ b/src/data/polynomial/splits.lean @@ -10,6 +10,9 @@ import data.polynomial.lifts /-! # Split polynomials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A polynomial `f : K[X]` splits over a field extension `L` of `K` if it is zero or all of its irreducible factors over `L` have degree `1`. diff --git a/src/data/set/list.lean b/src/data/set/list.lean index 78ea83204dfd8..afca35ed9df53 100644 --- a/src/data/set/list.lean +++ b/src/data/set/list.lean @@ -10,6 +10,9 @@ import data.fin.basic /-! # Lemmas about `list`s and `set.range` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove lemmas about range of some operations on lists. -/ diff --git a/src/data/set/pairwise/lattice.lean b/src/data/set/pairwise/lattice.lean index 46c97b39af121..1b27355c2f037 100644 --- a/src/data/set/pairwise/lattice.lean +++ b/src/data/set/pairwise/lattice.lean @@ -9,6 +9,9 @@ import data.set.pairwise.basic /-! # Relations holding pairwise +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove many facts about `pairwise` and the set lattice. -/ diff --git a/src/linear_algebra/affine_space/midpoint_zero.lean b/src/linear_algebra/affine_space/midpoint_zero.lean index a282c780150fe..a93b8c2d3bb5c 100644 --- a/src/linear_algebra/affine_space/midpoint_zero.lean +++ b/src/linear_algebra/affine_space/midpoint_zero.lean @@ -9,6 +9,9 @@ import linear_algebra.affine_space.midpoint /-! # Midpoint of a segment for characteristic zero +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We collect lemmas that require that the underlying ring has characteristic zero. ## Tags diff --git a/src/linear_algebra/affine_space/ordered.lean b/src/linear_algebra/affine_space/ordered.lean index a9ee6d0f2769a..6c746a980acb7 100644 --- a/src/linear_algebra/affine_space/ordered.lean +++ b/src/linear_algebra/affine_space/ordered.lean @@ -12,6 +12,9 @@ import tactic.field_simp /-! # Ordered modules as affine spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove some theorems about `slope` and `line_map` in the case when the module `E` acting on the codomain `PE` of a function is an ordered module over its domain `k`. We also prove inequalities that can be used to link convexity of a function on an interval to monotonicity of the diff --git a/src/linear_algebra/affine_space/pointwise.lean b/src/linear_algebra/affine_space/pointwise.lean index 8c861c3bc58d7..8dc1c4153421a 100644 --- a/src/linear_algebra/affine_space/pointwise.lean +++ b/src/linear_algebra/affine_space/pointwise.lean @@ -8,6 +8,9 @@ import linear_algebra.affine_space.affine_subspace /-! # Pointwise instances on `affine_subspace`s +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file provides the additive action `affine_subspace.pointwise_add_action` in the `pointwise` locale. diff --git a/src/linear_algebra/direct_sum/finsupp.lean b/src/linear_algebra/direct_sum/finsupp.lean index 4c7407e07feb9..5a369c076c87c 100644 --- a/src/linear_algebra/direct_sum/finsupp.lean +++ b/src/linear_algebra/direct_sum/finsupp.lean @@ -10,6 +10,9 @@ import linear_algebra.direct_sum.tensor_product /-! # Results on finitely supported functions. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The tensor product of ι →₀ M and κ →₀ N is linearly equivalent to (ι × κ) →₀ (M ⊗ N). -/ diff --git a/src/linear_algebra/direct_sum/tensor_product.lean b/src/linear_algebra/direct_sum/tensor_product.lean index 26e1f2acd5cf4..174207abfb52c 100644 --- a/src/linear_algebra/direct_sum/tensor_product.lean +++ b/src/linear_algebra/direct_sum/tensor_product.lean @@ -10,6 +10,9 @@ import algebra.direct_sum.module /-! # Tensor products of direct sums +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file shows that taking `tensor_product`s commutes with taking `direct_sum`s in both arguments. ## Main results diff --git a/src/model_theory/basic.lean b/src/model_theory/basic.lean index 5d97d4f852350..b9381bcfcd2e6 100644 --- a/src/model_theory/basic.lean +++ b/src/model_theory/basic.lean @@ -9,6 +9,9 @@ import set_theory.cardinal.basic /-! # Basics on First-Order Structures + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. This file defines first-order languages and structures in the style of the [Flypitch project](https://flypitch.github.io/), as well as several important maps between structures. diff --git a/src/number_theory/lucas_primality.lean b/src/number_theory/lucas_primality.lean index 8e565130d18e0..f2b85d805b377 100644 --- a/src/number_theory/lucas_primality.lean +++ b/src/number_theory/lucas_primality.lean @@ -11,6 +11,9 @@ import data.nat.totient /-! # The Lucas test for primes. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file implements the Lucas test for primes (not to be confused with the Lucas-Lehmer test for Mersenne primes). A number `a` witnesses that `n` is prime if `a` has order `n-1` in the multiplicative group of integers mod `n`. This is checked by verifying that `a^(n-1) = 1 (mod n)` diff --git a/src/order/upper_lower/locally_finite.lean b/src/order/upper_lower/locally_finite.lean index 4c8e171479c98..87855c6b90aa0 100644 --- a/src/order/upper_lower/locally_finite.lean +++ b/src/order/upper_lower/locally_finite.lean @@ -9,6 +9,9 @@ import order.upper_lower.basic /-! # Upper and lower sets in a locally finite order +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we characterise the interaction of `upper_set`/`lower_set` and `locally_finite_order`. -/ diff --git a/src/representation_theory/maschke.lean b/src/representation_theory/maschke.lean index 0eee7243865a4..6240b8f84dea3 100644 --- a/src/representation_theory/maschke.lean +++ b/src/representation_theory/maschke.lean @@ -10,6 +10,9 @@ import linear_algebra.basis /-! # Maschke's theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We prove **Maschke's theorem** for finite groups, in the formulation that every submodule of a `k[G]` module has a complement, when `k` is a field with `invertible (fintype.card G : k)`. diff --git a/src/ring_theory/adjoin/fg.lean b/src/ring_theory/adjoin/fg.lean index bcc8bfa719413..1c598c5365d47 100644 --- a/src/ring_theory/adjoin/fg.lean +++ b/src/ring_theory/adjoin/fg.lean @@ -10,6 +10,9 @@ import data.mv_polynomial.basic /-! # Adjoining elements to form subalgebras +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file develops the basic theory of finitely-generated subalgebras. ## Definitions diff --git a/src/ring_theory/localization/module.lean b/src/ring_theory/localization/module.lean index e9ceea7284331..efdab085b7ee2 100644 --- a/src/ring_theory/localization/module.lean +++ b/src/ring_theory/localization/module.lean @@ -10,6 +10,9 @@ import ring_theory.localization.integer /-! # Modules / vector spaces over localizations / fraction fields +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains some results about vector spaces over the field of fractions of a ring. ## Main results diff --git a/src/ring_theory/mv_polynomial/weighted_homogeneous.lean b/src/ring_theory/mv_polynomial/weighted_homogeneous.lean index 6d6edcc4d1f1c..5e16ce5bad6d8 100644 --- a/src/ring_theory/mv_polynomial/weighted_homogeneous.lean +++ b/src/ring_theory/mv_polynomial/weighted_homogeneous.lean @@ -9,6 +9,9 @@ import data.mv_polynomial.variables /-! # Weighted homogeneous polynomials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + It is possible to assign weights (in a commutative additive monoid `M`) to the variables of a multivariate polynomial ring, so that monomials of the ring then have a weighted degree with respect to the weights of the variables. The weights are represented by a function `w : σ → M`, diff --git a/src/ring_theory/polynomial/basic.lean b/src/ring_theory/polynomial/basic.lean index e31e66cc6bad7..f287919dda4e2 100644 --- a/src/ring_theory/polynomial/basic.lean +++ b/src/ring_theory/polynomial/basic.lean @@ -14,6 +14,9 @@ import ring_theory.unique_factorization_domain /-! # Ring-theoretic supplement of data.polynomial. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main results * `mv_polynomial.is_domain`: If a ring is an integral domain, then so is its polynomial ring over finitely many variables. diff --git a/src/ring_theory/polynomial/vieta.lean b/src/ring_theory/polynomial/vieta.lean index 2ae9290898d00..d3687f762f0b2 100644 --- a/src/ring_theory/polynomial/vieta.lean +++ b/src/ring_theory/polynomial/vieta.lean @@ -9,6 +9,9 @@ import ring_theory.mv_polynomial.symmetric /-! # Vieta's Formula +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The main result is `multiset.prod_X_add_C_eq_sum_esymm`, which shows that the product of linear terms `X + λ` with `λ` in a `multiset s` is equal to a linear combination of the symmetric functions `esymm s`. diff --git a/src/topology/algebra/polynomial.lean b/src/topology/algebra/polynomial.lean index 4467295b0d38c..41202aaff5cda 100644 --- a/src/topology/algebra/polynomial.lean +++ b/src/topology/algebra/polynomial.lean @@ -12,6 +12,9 @@ import analysis.normed.field.basic /-! # Polynomials and limits +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove the following lemmas. * `polynomial.continuous_eval₂: `polynomial.eval₂` defines a continuous function. From 47a1a73351de8dd6c8d3d32b569c8e434b03ca47 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sun, 2 Apr 2023 23:06:01 +0000 Subject: [PATCH 0760/1291] =?UTF-8?q?feat(data/{int,nat}/modeq):=20`a/c=20?= =?UTF-8?q?=E2=89=A1=20b/c=20mod=20m/c=20=E2=86=92=20a=20=E2=89=A1=20b=20m?= =?UTF-8?q?od=20m`=20(#18666)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Also prove `-a ≡ -b [ZMOD n] ↔ a ≡ b [ZMOD n]`, `a ≡ b [ZMOD -n] ↔ a ≡ b [ZMOD n]`, generalise `int.modeq.mul_left'`/`int.modeq.mul_right'`, and rename * `int.gcd_pos_of_non_zero_left` → `int.gcd_pos_of_ne_zero_left` * `int.gcd_pos_of_non_zero_right` → `int.gcd_pos_of_ne_zero_right` * `eq_iff_modeq_int`, `char_p.int_coe_eq_int_coe_iff` → `char_p.int_cast_eq_int_cast` (they were duplicates) --- src/algebra/char_p/basic.lean | 19 +++++----- src/algebra/ring/divisibility.lean | 2 + src/data/int/gcd.lean | 10 ++--- src/data/int/modeq.lean | 60 ++++++++++++++++++++++++------ src/data/nat/modeq.lean | 4 ++ src/data/nat/order/lemmas.lean | 9 ++++- src/data/zmod/basic.lean | 2 +- 7 files changed, 78 insertions(+), 28 deletions(-) diff --git a/src/algebra/char_p/basic.lean b/src/algebra/char_p/basic.lean index dae05cea243f2..194a45351e265 100644 --- a/src/algebra/char_p/basic.lean +++ b/src/algebra/char_p/basic.lean @@ -121,10 +121,16 @@ begin rw [int.cast_coe_nat, char_p.cast_eq_zero_iff R p, int.coe_nat_dvd] } end -lemma char_p.int_coe_eq_int_coe_iff [add_group_with_one R] (p : ℕ) [char_p R p] (a b : ℤ) : - (a : R) = (b : R) ↔ a ≡ b [ZMOD p] := -by rw [eq_comm, ←sub_eq_zero, ←int.cast_sub, - char_p.int_cast_eq_zero_iff R p, int.modeq_iff_dvd] +lemma char_p.int_cast_eq_int_cast [add_group_with_one R] (p : ℕ) [char_p R p] {a b : ℤ} : + (a : R) = b ↔ a ≡ b [ZMOD p] := +by rw [eq_comm, ←sub_eq_zero, ←int.cast_sub, char_p.int_cast_eq_zero_iff R p, int.modeq_iff_dvd] + +lemma char_p.nat_cast_eq_nat_cast [add_group_with_one R] (p : ℕ) [char_p R p] {a b : ℕ} : + (a : R) = b ↔ a ≡ b [MOD p] := +begin + rw [←int.cast_coe_nat, ←int.cast_coe_nat b], + exact (char_p.int_cast_eq_int_cast _ _).trans int.coe_nat_modeq_iff, +end theorem char_p.eq [add_monoid_with_one R] {p q : ℕ} (c1 : char_p R p) (c2 : char_p R q) : p = q := @@ -242,11 +248,6 @@ theorem sub_pow_char_pow [comm_ring R] {p : ℕ} [fact p.prime] (x - y) ^ (p ^ n) = x ^ (p ^ n) - y ^ (p ^ n) := sub_pow_char_pow_of_commute _ _ _ (commute.all _ _) -lemma eq_iff_modeq_int [ring R] (p : ℕ) [char_p R p] (a b : ℤ) : - (a : R) = b ↔ a ≡ b [ZMOD p] := -by rw [eq_comm, ←sub_eq_zero, ←int.cast_sub, - char_p.int_cast_eq_zero_iff R p, int.modeq_iff_dvd] - lemma char_p.neg_one_ne_one [ring R] (p : ℕ) [char_p R p] [fact (2 < p)] : (-1 : R) ≠ (1 : R) := begin diff --git a/src/algebra/ring/divisibility.lean b/src/algebra/ring/divisibility.lean index 21dcb248c7b1c..00ee02cd96226 100644 --- a/src/algebra/ring/divisibility.lean +++ b/src/algebra/ring/divisibility.lean @@ -96,6 +96,8 @@ begin exact eq_add_of_sub_eq rfl } end +lemma dvd_sub_comm : a ∣ b - c ↔ a ∣ c - b := by rw [←dvd_neg, neg_sub] + end non_unital_ring section ring diff --git a/src/data/int/gcd.lean b/src/data/int/gcd.lean index 50f50b40dd6b5..e05c44e65c5e4 100644 --- a/src/data/int/gcd.lean +++ b/src/data/int/gcd.lean @@ -214,11 +214,11 @@ by { rw [int.gcd, int.gcd, nat_abs_mul, nat_abs_mul], apply nat.gcd_mul_left } theorem gcd_mul_right (i j k : ℤ) : gcd (i * j) (k * j) = gcd i k * nat_abs j := by { rw [int.gcd, int.gcd, nat_abs_mul, nat_abs_mul], apply nat.gcd_mul_right } -theorem gcd_pos_of_non_zero_left {i : ℤ} (j : ℤ) (i_non_zero : i ≠ 0) : 0 < gcd i j := -nat.gcd_pos_of_pos_left (nat_abs j) (nat_abs_pos_of_ne_zero i_non_zero) +theorem gcd_pos_of_ne_zero_left {i : ℤ} (j : ℤ) (hi : i ≠ 0) : 0 < gcd i j := +nat.gcd_pos_of_pos_left _ $ nat_abs_pos_of_ne_zero hi -theorem gcd_pos_of_non_zero_right (i : ℤ) {j : ℤ} (j_non_zero : j ≠ 0) : 0 < gcd i j := -nat.gcd_pos_of_pos_right (nat_abs i) (nat_abs_pos_of_ne_zero j_non_zero) +theorem gcd_pos_of_ne_zero_right (i : ℤ) {j : ℤ} (hj : j ≠ 0) : 0 < gcd i j := +nat.gcd_pos_of_pos_right _ $ nat_abs_pos_of_ne_zero hj theorem gcd_eq_zero_iff {i j : ℤ} : gcd i j = 0 ↔ i = 0 ∧ j = 0 := begin @@ -339,7 +339,7 @@ theorem gcd_least_linear {a b : ℤ} (ha : a ≠ 0) : begin simp_rw ←gcd_dvd_iff, split, - { simpa [and_true, dvd_refl, set.mem_set_of_eq] using gcd_pos_of_non_zero_left b ha }, + { simpa [and_true, dvd_refl, set.mem_set_of_eq] using gcd_pos_of_ne_zero_left b ha }, { simp only [lower_bounds, and_imp, set.mem_set_of_eq], exact λ n hn_pos hn, nat.le_of_dvd hn_pos hn }, end diff --git a/src/data/int/modeq.lean b/src/data/int/modeq.lean index 493111c330de6..0476b3250506c 100644 --- a/src/data/int/modeq.lean +++ b/src/data/int/modeq.lean @@ -45,8 +45,12 @@ instance : is_refl _ (modeq n) := ⟨modeq.refl⟩ @[trans] protected theorem trans : a ≡ b [ZMOD n] → b ≡ c [ZMOD n] → a ≡ c [ZMOD n] := eq.trans +protected lemma eq : a ≡ b [ZMOD n] → a % n = b % n := id + end modeq +lemma modeq_comm : a ≡ b [ZMOD n] ↔ b ≡ a [ZMOD n] := ⟨modeq.symm, modeq.symm⟩ + lemma coe_nat_modeq_iff {a b n : ℕ} : a ≡ b [ZMOD n] ↔ a ≡ b [MOD n] := by unfold modeq nat.modeq; rw ← int.coe_nat_eq_coe_nat_iff; simp [coe_nat_mod] @@ -70,19 +74,27 @@ alias modeq_iff_dvd ↔ modeq.dvd modeq_of_dvd theorem mod_modeq (a n) : a % n ≡ a [ZMOD n] := mod_mod _ _ +@[simp] lemma neg_modeq_neg : -a ≡ -b [ZMOD n] ↔ a ≡ b [ZMOD n] := +by simp [modeq_iff_dvd, dvd_sub_comm] + +@[simp] lemma modeq_neg : a ≡ b [ZMOD -n] ↔ a ≡ b [ZMOD n] := by simp [modeq_iff_dvd] + namespace modeq protected lemma of_dvd (d : m ∣ n) (h : a ≡ b [ZMOD n]) : a ≡ b [ZMOD m] := modeq_of_dvd $ d.trans h.dvd -protected theorem mul_left' (hc : 0 ≤ c) (h : a ≡ b [ZMOD n]) : c * a ≡ c * b [ZMOD (c * n)] := -or.cases_on hc.lt_or_eq (λ hc, - by unfold modeq; - simp [mul_mod_mul_of_pos hc, (show _ = _, from h)] ) -(λ hc, by simp [hc.symm]) +protected theorem mul_left' (h : a ≡ b [ZMOD n]) : c * a ≡ c * b [ZMOD (c * n)] := +begin + obtain hc | rfl | hc := lt_trichotomy c 0, + { rw [←neg_modeq_neg, ←modeq_neg, ←neg_mul, ←neg_mul, ←neg_mul], + simp only [modeq, mul_mod_mul_of_pos (neg_pos.2 hc), h.eq] }, + { simp }, + { simp only [modeq, mul_mod_mul_of_pos hc, h.eq] } +end -protected theorem mul_right' (hc : 0 ≤ c) (h : a ≡ b [ZMOD n]) : a * c ≡ b * c [ZMOD (n * c)] := -by rw [mul_comm a, mul_comm b, mul_comm n]; exact h.mul_left' hc +protected theorem mul_right' (h : a ≡ b [ZMOD n]) : a * c ≡ b * c [ZMOD (n * c)] := +by rw [mul_comm a, mul_comm b, mul_comm n]; exact h.mul_left' protected theorem add (h₁ : a ≡ b [ZMOD n]) (h₂ : c ≡ d [ZMOD n]) : a + c ≡ b + d [ZMOD n] := modeq_iff_dvd.2 $ by { convert dvd_add h₁.dvd h₂.dvd, ring } @@ -122,13 +134,10 @@ protected theorem sub_right (c : ℤ) (h : a ≡ b [ZMOD n]) : a - c ≡ b - c [ h.sub modeq.rfl protected theorem mul_left (c : ℤ) (h : a ≡ b [ZMOD n]) : c * a ≡ c * b [ZMOD n] := -or.cases_on (le_total 0 c) -(λ hc, (h.mul_left' hc).of_dvd (dvd_mul_left _ _) ) -(λ hc, by rw [← neg_neg c, neg_mul, neg_mul _ b]; - exact ((h.mul_left' $ neg_nonneg.2 hc).of_dvd (dvd_mul_left _ _)).neg) +h.mul_left'.of_dvd $ dvd_mul_left _ _ protected theorem mul_right (c : ℤ) (h : a ≡ b [ZMOD n]) : a * c ≡ b * c [ZMOD n] := -by { rw [mul_comm a, mul_comm b], exact h.mul_left c } +h.mul_right'.of_dvd $ dvd_mul_right _ _ protected theorem mul (h₁ : a ≡ b [ZMOD n]) (h₂ : c ≡ d [ZMOD n]) : a * c ≡ b * d [ZMOD n] := (h₂.mul_left _).trans (h₁.mul_right _) @@ -146,6 +155,28 @@ by rw [modeq_iff_dvd] at *; exact (dvd_mul_left n m).trans h theorem of_mul_right (m : ℤ) : a ≡ b [ZMOD n * m] → a ≡ b [ZMOD n] := mul_comm m n ▸ of_mul_left _ +/-- To cancel a common factor `c` from a `modeq` we must divide the modulus `m` by `gcd m c`. -/ +lemma cancel_right_div_gcd (hm : 0 < m) (h : a * c ≡ b * c [ZMOD m]) : a ≡ b [ZMOD m / gcd m c] := +begin + let d := gcd m c, + have hmd := gcd_dvd_left m c, + have hcd := gcd_dvd_right m c, + rw modeq_iff_dvd at ⊢ h, + refine int.dvd_of_dvd_mul_right_of_gcd_one _ _, + show m / d ∣ c / d * (b - a), + { rw [mul_comm, ←int.mul_div_assoc (b - a) hcd, sub_mul], + exact int.div_dvd_div hmd h }, + { rw [gcd_div hmd hcd, nat_abs_of_nat, nat.div_self (gcd_pos_of_ne_zero_left c hm.ne')] } +end + +/-- To cancel a common factor `c` from a `modeq` we must divide the modulus `m` by `gcd m c`. -/ +lemma cancel_left_div_gcd (hm : 0 < m) (h : c * a ≡ c * b [ZMOD m]) : a ≡ b [ZMOD m / gcd m c] := +cancel_right_div_gcd hm $ by simpa [mul_comm] using h + +lemma of_div (h : a / c ≡ b / c [ZMOD m / c]) (ha : c ∣ a) (ha : c ∣ b) (ha : c ∣ m) : + a ≡ b [ZMOD m] := +by convert h.mul_left'; rwa int.mul_div_cancel' + end modeq theorem modeq_one : a ≡ b [ZMOD 1] := modeq_of_dvd (one_dvd _) @@ -153,6 +184,11 @@ theorem modeq_one : a ≡ b [ZMOD 1] := modeq_of_dvd (one_dvd _) lemma modeq_sub (a b : ℤ) : a ≡ b [ZMOD a - b] := (modeq_of_dvd dvd_rfl).symm +@[simp] lemma modeq_zero_iff : a ≡ b [ZMOD 0] ↔ a = b := by rw [modeq, mod_zero, mod_zero] + +@[simp] lemma add_modeq_left : n + a ≡ a [ZMOD n] := modeq.symm $ modeq_iff_dvd.2 $ by simp +@[simp] lemma add_modeq_right : a + n ≡ a [ZMOD n] := modeq.symm $ modeq_iff_dvd.2 $ by simp + lemma modeq_and_modeq_iff_modeq_mul {a b m n : ℤ} (hmn : m.nat_abs.coprime n.nat_abs) : a ≡ b [ZMOD m] ∧ a ≡ b [ZMOD n] ↔ (a ≡ b [ZMOD m * n]) := ⟨λ h, begin diff --git a/src/data/nat/modeq.lean b/src/data/nat/modeq.lean index e3340f3d69167..3ac75bbcc4506 100644 --- a/src/data/nat/modeq.lean +++ b/src/data/nat/modeq.lean @@ -160,6 +160,10 @@ by { rw [modeq_iff_dvd] at *, exact (dvd_mul_left (n : ℤ) (m : ℤ)).trans h } For cancelling right multiplication on both sides of the `≡`, see `nat.modeq.mul_right_cancel'`. -/ theorem of_mul_right (m : ℕ) : a ≡ b [MOD n * m] → a ≡ b [MOD n] := mul_comm m n ▸ of_mul_left _ +lemma of_div (h : a / c ≡ b / c [MOD m / c]) (ha : c ∣ a) (ha : c ∣ b) (ha : c ∣ m) : + a ≡ b [MOD m] := +by convert h.mul_left' c; rwa nat.mul_div_cancel' + end modeq lemma modeq_sub (h : b ≤ a) : a ≡ b [MOD a - b] := (modeq_of_dvd $ by rw [int.coe_nat_sub h]).symm diff --git a/src/data/nat/order/lemmas.lean b/src/data/nat/order/lemmas.lean index 29c162696faec..95e76cfe69ee5 100644 --- a/src/data/nat/order/lemmas.lean +++ b/src/data/nat/order/lemmas.lean @@ -21,7 +21,7 @@ mathlib4. After `data.rat.order` has been ported, please feel free to reorganize universes u v -variables {m n k : ℕ} +variables {a b m n k : ℕ} namespace nat /-! ### Sets -/ @@ -183,6 +183,13 @@ lemma eq_zero_of_dvd_of_lt {a b : ℕ} (w : a ∣ b) (h : b < a) : b = 0 := nat.eq_zero_of_dvd_of_div_eq_zero w ((nat.div_eq_zero_iff (lt_of_le_of_lt (zero_le b) h)).elim_right h) +lemma le_of_lt_add_of_dvd (h : a < b + n) : n ∣ a → n ∣ b → a ≤ b := +begin + rintro ⟨a, rfl⟩ ⟨b, rfl⟩, + rw ←mul_add_one at h, + exact mul_le_mul_left' (lt_succ_iff.1 $ lt_of_mul_lt_mul_left h bot_le) _, +end + @[simp] lemma mod_div_self (m n : ℕ) : m % n / n = 0 := begin cases n, diff --git a/src/data/zmod/basic.lean b/src/data/zmod/basic.lean index f751ced32c8f4..bf03ebdeae0e2 100644 --- a/src/data/zmod/basic.lean +++ b/src/data/zmod/basic.lean @@ -383,7 +383,7 @@ end universal_property lemma int_coe_eq_int_coe_iff (a b : ℤ) (c : ℕ) : (a : zmod c) = (b : zmod c) ↔ a ≡ b [ZMOD c] := -char_p.int_coe_eq_int_coe_iff (zmod c) c a b +char_p.int_cast_eq_int_cast (zmod c) c lemma int_coe_eq_int_coe_iff' (a b : ℤ) (c : ℕ) : (a : zmod c) = (b : zmod c) ↔ a % c = b % c := From 94eaaaa6064d32e98cf838789144cf5318c37cf0 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Mon, 3 Apr 2023 05:43:55 +0000 Subject: [PATCH 0761/1291] chore(*): add mathlib4 synchronization comments (#18721) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `category_theory.monad.kleisli` * `data.matrix.basic` * `ring_theory.adjoin.tower` --- src/category_theory/monad/kleisli.lean | 3 +++ src/data/matrix/basic.lean | 3 +++ src/ring_theory/adjoin/tower.lean | 3 +++ 3 files changed, 9 insertions(+) diff --git a/src/category_theory/monad/kleisli.lean b/src/category_theory/monad/kleisli.lean index 138d9ba68702b..0f3bff8e8f12b 100644 --- a/src/category_theory/monad/kleisli.lean +++ b/src/category_theory/monad/kleisli.lean @@ -9,6 +9,9 @@ import category_theory.monad.basic /-! # Kleisli category on a (co)monad +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the Kleisli category on a monad `(T, η_ T, μ_ T)` as well as the co-Kleisli category on a comonad `(U, ε_ U, δ_ U)`. It also defines the Kleisli adjunction which gives rise to the monad `(T, η_ T, μ_ T)` as well as the co-Kleisli adjunction which gives rise to the comonad diff --git a/src/data/matrix/basic.lean b/src/data/matrix/basic.lean index 76c5cb08b2861..65388afef5547 100644 --- a/src/data/matrix/basic.lean +++ b/src/data/matrix/basic.lean @@ -18,6 +18,9 @@ import data.fintype.big_operators /-! # Matrices +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines basic properties of matrices. Matrices with rows indexed by `m`, columns indexed by `n`, and entries of type `α` are represented diff --git a/src/ring_theory/adjoin/tower.lean b/src/ring_theory/adjoin/tower.lean index ee0cf59cea822..eb67e7e843c1f 100644 --- a/src/ring_theory/adjoin/tower.lean +++ b/src/ring_theory/adjoin/tower.lean @@ -8,6 +8,9 @@ import ring_theory.adjoin.fg /-! # Adjoining elements and being finitely generated in an algebra tower +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main results * `algebra.fg_trans'`: if `S` is finitely generated as `R`-algebra and `A` as `S`-algebra, From 7c2ce0c2da15516b4e65d0c9e254bb6dc93abd1f Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Mon, 3 Apr 2023 08:02:26 +0000 Subject: [PATCH 0762/1291] refactor(set_theory/cardinal/basic): redefine `cardinal.is_limit` in terms of `is_succ_limit` (#18523) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We redefine `cardinal.is_limit x` as `x ≠ 0 ∧ is_succ_limit x`. This will allow us to make use of the extensive `is_succ_limit` API in the future. --- src/set_theory/cardinal/basic.lean | 35 ++++++++++++++++++++++++- src/set_theory/cardinal/cofinality.lean | 34 +++++------------------- src/set_theory/cardinal/ordinal.lean | 8 +++--- 3 files changed, 45 insertions(+), 32 deletions(-) diff --git a/src/set_theory/cardinal/basic.lean b/src/set_theory/cardinal/basic.lean index 5022c5826899b..8c1ac0bae9315 100644 --- a/src/set_theory/cardinal/basic.lean +++ b/src/set_theory/cardinal/basic.lean @@ -9,7 +9,7 @@ import data.nat.part_enat import data.set.countable import logic.small.basic import order.conditionally_complete_lattice.basic -import order.succ_pred.basic +import order.succ_pred.limit import set_theory.cardinal.schroeder_bernstein import tactic.positivity @@ -30,6 +30,7 @@ We define cardinal numbers as a quotient of types under the equivalence relation * Multiplication `c₁ * c₂` is defined by `cardinal.mul_def : #α * #β = #(α × β)`. * The order `c₁ ≤ c₂` is defined by `cardinal.le_def α β : #α ≤ #β ↔ nonempty (α ↪ β)`. * Exponentiation `c₁ ^ c₂` is defined by `cardinal.power_def α β : #α ^ #β = #(β → α)`. +* `cardinal.is_limit c` means that `c` is a (weak) limit cardinal: `c ≠ 0 ∧ ∀ x < c, succ x < c`. * `cardinal.aleph_0` or `ℵ₀` is the cardinality of `ℕ`. This definition is universe polymorphic: `cardinal.aleph_0.{u} : cardinal.{u}` (contrast with `ℕ : Type`, which lives in a specific universe). In some cases the universe level has to be given explicitly. @@ -560,6 +561,20 @@ instance : succ_order cardinal := succ_order.of_succ_le_iff (λ c, Inf {c' | c < c'}) (λ a b, ⟨lt_of_lt_of_le $ Inf_mem $ exists_gt a, cInf_le'⟩) +/-- A cardinal is a limit if it is not zero or a successor cardinal. Note that `ℵ₀` is a limit + cardinal by this definition, but `0` isn't. + + Use `is_succ_limit` if you want to include the `c = 0` case. -/ +def is_limit (c : cardinal) : Prop := c ≠ 0 ∧ is_succ_limit c + +protected theorem is_limit.ne_zero {c} (h : is_limit c) : c ≠ 0 := h.1 + +protected theorem is_limit.is_succ_limit {c} (h : is_limit c) : is_succ_limit c := h.2 + +theorem is_limit.succ_lt {x c} (h : is_limit c) : x < c → succ x < c := h.is_succ_limit.succ_lt + +theorem is_succ_limit_zero : is_succ_limit (0 : cardinal) := is_succ_limit_bot + theorem succ_def (c : cardinal) : succ c = Inf {c' | c < c'} := rfl theorem add_one_le_succ (c : cardinal.{u}) : c + 1 ≤ succ c := @@ -1000,6 +1015,24 @@ lt_aleph_0_iff_finite.trans finite_coe_iff alias lt_aleph_0_iff_set_finite ↔ _ _root_.set.finite.lt_aleph_0 +theorem is_succ_limit_aleph_0 : is_succ_limit ℵ₀ := +is_succ_limit_of_succ_lt $ λ a ha, begin + rcases lt_aleph_0.1 ha with ⟨n, rfl⟩, + rw ←nat_succ, + apply nat_lt_aleph_0 +end + +theorem is_limit_aleph_0 : is_limit ℵ₀ := ⟨aleph_0_ne_zero, is_succ_limit_aleph_0⟩ + +theorem is_limit.aleph_0_le {c : cardinal} (h : is_limit c) : ℵ₀ ≤ c := +begin + by_contra' h', + rcases lt_aleph_0.1 h' with ⟨_ | n, rfl⟩, + { exact h.ne_zero.irrefl }, + { rw nat_succ at h, + exact not_is_succ_limit_succ _ h.is_succ_limit } +end + @[simp] theorem lt_aleph_0_iff_subtype_finite {p : α → Prop} : #{x // p x} < ℵ₀ ↔ {x | p x}.finite := lt_aleph_0_iff_set_finite diff --git a/src/set_theory/cardinal/cofinality.lean b/src/set_theory/cardinal/cofinality.lean index 8fd6d2b3dd19a..9d59e9cda57a7 100644 --- a/src/set_theory/cardinal/cofinality.lean +++ b/src/set_theory/cardinal/cofinality.lean @@ -20,7 +20,6 @@ This file contains the definition of cofinality of an ordinal number and regular * `ordinal.cof o` is the cofinality of the ordinal `o`. If `o` is the order type of the relation `<` on `α`, then `o.cof` is the smallest cardinality of a subset `s` of α that is *cofinal* in `α`, i.e. `∀ x : α, ∃ y ∈ s, ¬ y < x`. -* `cardinal.is_limit c` means that `c` is a (weak) limit cardinal: `c ≠ 0 ∧ ∀ x < c, succ x < c`. * `cardinal.is_strong_limit c` means that `c` is a strong limit cardinal: `c ≠ 0 ∧ ∀ x < c, 2 ^ x < c`. * `cardinal.is_regular c` means that `c` is a regular cardinal: `ℵ₀ ≤ c ∧ c.ord.cof = c`. @@ -715,25 +714,6 @@ open ordinal local infixr (name := cardinal.pow) ^ := @pow cardinal.{u} cardinal cardinal.has_pow -/-- A cardinal is a limit if it is not zero or a successor - cardinal. Note that `ℵ₀` is a limit cardinal by this definition. -/ -def is_limit (c : cardinal) : Prop := -c ≠ 0 ∧ ∀ x < c, succ x < c - -theorem is_limit.ne_zero {c} (h : is_limit c) : c ≠ 0 := -h.1 - -theorem is_limit.succ_lt {x c} (h : is_limit c) : x < c → succ x < c := -h.2 x - -theorem is_limit.aleph_0_le {c} (h : is_limit c) : ℵ₀ ≤ c := -begin - by_contra' h', - rcases lt_aleph_0.1 h' with ⟨_ | n, rfl⟩, - { exact h.1.irrefl }, - { simpa using h.2 n } -end - /-- A cardinal is a strong limit if it is not zero and it is closed under powersets. Note that `ℵ₀` is a strong limit by this definition. -/ def is_strong_limit (c : cardinal) : Prop := @@ -751,23 +731,23 @@ theorem is_strong_limit_aleph_0 : is_strong_limit ℵ₀ := exact_mod_cast nat_lt_aleph_0 (pow 2 n) end⟩ -theorem is_strong_limit.is_limit {c} (H : is_strong_limit c) : is_limit c := -⟨H.1, λ x h, (succ_le_of_lt $ cantor x).trans_lt (H.2 _ h)⟩ +protected theorem is_strong_limit.is_succ_limit {c} (H : is_strong_limit c) : is_succ_limit c := +is_succ_limit_of_succ_lt $ λ x h, (succ_le_of_lt $ cantor x).trans_lt (H.two_power_lt h) -theorem is_limit_aleph_0 : is_limit ℵ₀ := -is_strong_limit_aleph_0.is_limit +theorem is_strong_limit.is_limit {c} (H : is_strong_limit c) : is_limit c := +⟨H.ne_zero, H.is_succ_limit⟩ -theorem is_strong_limit_beth {o : ordinal} (H : ∀ a < o, succ a < o) : is_strong_limit (beth o) := +theorem is_strong_limit_beth {o : ordinal} (H : is_succ_limit o) : is_strong_limit (beth o) := begin rcases eq_or_ne o 0 with rfl | h, { rw beth_zero, exact is_strong_limit_aleph_0 }, { refine ⟨beth_ne_zero o, λ a ha, _⟩, - rw beth_limit ⟨h, H⟩ at ha, + rw beth_limit ⟨h, is_succ_limit_iff_succ_lt.1 H⟩ at ha, rcases exists_lt_of_lt_csupr' ha with ⟨⟨i, hi⟩, ha⟩, have := power_le_power_left two_ne_zero ha.le, rw ←beth_succ at this, - exact this.trans_lt (beth_lt.2 (H i hi)) } + exact this.trans_lt (beth_lt.2 (H.succ_lt hi)) } end theorem mk_bounded_subset {α : Type*} (h : ∀ x < #α, 2 ^ x < #α) {r : α → α → Prop} diff --git a/src/set_theory/cardinal/ordinal.lean b/src/set_theory/cardinal/ordinal.lean index 908116dd92b01..655c3a267b92c 100644 --- a/src/set_theory/cardinal/ordinal.lean +++ b/src/set_theory/cardinal/ordinal.lean @@ -182,7 +182,7 @@ theorem aleph'_le_of_limit {o : ordinal} (l : o.is_limit) {c} : exact h _ h' end⟩ -theorem aleph'_limit {o : ordinal} (ho : is_limit o) : aleph' o = ⨆ a : Iio o, aleph' a := +theorem aleph'_limit {o : ordinal} (ho : o.is_limit) : aleph' o = ⨆ a : Iio o, aleph' a := begin refine le_antisymm _ (csupr_le' (λ i, aleph'_le.2 (le_of_lt i.2))), rw aleph'_le_of_limit ho, @@ -223,7 +223,7 @@ by rw [aleph, add_succ, aleph'_succ, aleph] @[simp] theorem aleph_zero : aleph 0 = ℵ₀ := by rw [aleph, add_zero, aleph'_omega] -theorem aleph_limit {o : ordinal} (ho : is_limit o) : aleph o = ⨆ a : Iio o, aleph a := +theorem aleph_limit {o : ordinal} (ho : o.is_limit) : aleph o = ⨆ a : Iio o, aleph a := begin apply le_antisymm _ (csupr_le' _), { rw [aleph, aleph'_limit (ho.add _)], @@ -261,7 +261,7 @@ begin exact λ h, (ord_injective h).not_gt (aleph_pos o) end -theorem ord_aleph_is_limit (o : ordinal) : is_limit (aleph o).ord := +theorem ord_aleph_is_limit (o : ordinal) : (aleph o).ord.is_limit := ord_is_limit $ aleph_0_le_aleph _ instance (o : ordinal) : no_max_order (aleph o).ord.out.α := @@ -347,7 +347,7 @@ limit_rec_on_zero _ _ _ @[simp] theorem beth_succ (o : ordinal) : beth (succ o) = 2 ^ beth o := limit_rec_on_succ _ _ _ _ -theorem beth_limit {o : ordinal} : is_limit o → beth o = ⨆ a : Iio o, beth a := +theorem beth_limit {o : ordinal} : o.is_limit → beth o = ⨆ a : Iio o, beth a := limit_rec_on_limit _ _ _ _ theorem beth_strict_mono : strict_mono beth := From 59628387770d82eb6f6dd7b7107308aa2509ec95 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 3 Apr 2023 10:26:19 +0000 Subject: [PATCH 0763/1291] refactor(linear_algebra): make `module.free` available inside `linear_algebra/dimension` (#18717) Many of the results here should generalize from `division_ring K` to `module.free K V`, though this PR doesn't bother itself with making them. This PR just flips around the imports and moves the lemmas that can't stay in the old home, and makes no attempt to actually make this generalization. This also removes some duplicates: * `lemma equiv_of_dim_eq_lift_dim` duplicates `nonempty_linear_equiv_of_lift_dim_eq` * `def equiv_of_dim_eq_dim` duplicates `linear_equiv.of_dim_eq` A few downstream files now need to directly import `linear_algebra.dimension`, which previously was implied transitively. --- src/field_theory/mv_polynomial.lean | 2 + src/linear_algebra/dimension.lean | 26 ++++++++ src/linear_algebra/finsupp_vector_space.lean | 64 +------------------ src/linear_algebra/free_algebra.lean | 1 + .../free_module/finite/basic.lean | 1 + src/linear_algebra/free_module/rank.lean | 3 +- 6 files changed, 33 insertions(+), 64 deletions(-) diff --git a/src/field_theory/mv_polynomial.lean b/src/field_theory/mv_polynomial.lean index fbe21e471a266..d55696ff1367d 100644 --- a/src/field_theory/mv_polynomial.lean +++ b/src/field_theory/mv_polynomial.lean @@ -5,6 +5,8 @@ Authors: Johannes Hölzl -/ import data.mv_polynomial.comm_ring +import linear_algebra.dimension +import ring_theory.ideal.quotient import ring_theory.mv_polynomial.basic /-! diff --git a/src/linear_algebra/dimension.lean b/src/linear_algebra/dimension.lean index 1dc84fa2970ef..b01c0b30f9395 100644 --- a/src/linear_algebra/dimension.lean +++ b/src/linear_algebra/dimension.lean @@ -5,6 +5,7 @@ Authors: Mario Carneiro, Johannes Hölzl, Sander Dahmen, Scott Morrison -/ import algebra.module.big_operators import linear_algebra.dfinsupp +import linear_algebra.free_module.basic import linear_algebra.invariant_basis_number import linear_algebra.isomorphisms import linear_algebra.std_basis @@ -1354,3 +1355,28 @@ end end division_ring end module + +section field +variables [field K] [add_comm_group V] [module K V] + +lemma finsupp.dim_eq {ι : Type v} : module.rank K (ι →₀ V) = #ι * module.rank K V := +begin + let bs := basis.of_vector_space K V, + rw [← bs.mk_eq_dim'', ← (finsupp.basis (λa:ι, bs)).mk_eq_dim'', + cardinal.mk_sigma, cardinal.sum_const'] +end + +-- TODO: merge with the `finrank` content +/-- An `n`-dimensional `K`-vector space is equivalent to `fin n → K`. -/ +def fin_dim_vectorspace_equiv (n : ℕ) + (hn : (module.rank K V) = n) : V ≃ₗ[K] (fin n → K) := +begin + have : cardinal.lift.{u} (n : cardinal.{v}) = cardinal.lift.{v} (n : cardinal.{u}), + by simp, + have hn := cardinal.lift_inj.{v u}.2 hn, + rw this at hn, + rw ←@dim_fin_fun K _ n at hn, + exact classical.choice (nonempty_linear_equiv_of_lift_dim_eq hn), +end + +end field diff --git a/src/linear_algebra/finsupp_vector_space.lean b/src/linear_algebra/finsupp_vector_space.lean index 143572dea700a..82d263f0d7a13 100644 --- a/src/linear_algebra/finsupp_vector_space.lean +++ b/src/linear_algebra/finsupp_vector_space.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl -/ -import linear_algebra.dimension import linear_algebra.std_basis /-! @@ -13,11 +12,6 @@ import linear_algebra.std_basis This file contains results on the `R`-module structure on functions of finite support from a type `ι` to an `R`-module `M`, in particular in the case that `R` is a field. -Furthermore, it contains some facts about isomorphisms of vector spaces from equality of dimension. - -## TODO - -Move the second half of this file to more appropriate other files. -/ noncomputable theory @@ -122,65 +116,9 @@ funext $ λ i, basis.apply_eq_iff.mpr rfl end semiring -section dim -variables {K : Type u} {V : Type v} {ι : Type v} -variables [field K] [add_comm_group V] [module K V] - -lemma dim_eq : module.rank K (ι →₀ V) = #ι * module.rank K V := -begin - let bs := basis.of_vector_space K V, - rw [← bs.mk_eq_dim'', ← (finsupp.basis (λa:ι, bs)).mk_eq_dim'', - cardinal.mk_sigma, cardinal.sum_const'] -end - -end dim - end finsupp -section module -variables {K : Type u} {V V₁ V₂ : Type v} {V' : Type w} -variables [field K] -variables [add_comm_group V] [module K V] -variables [add_comm_group V₁] [module K V₁] -variables [add_comm_group V₂] [module K V₂] -variables [add_comm_group V'] [module K V'] - -open module - -lemma equiv_of_dim_eq_lift_dim - (h : cardinal.lift.{w} (module.rank K V) = cardinal.lift.{v} (module.rank K V')) : - nonempty (V ≃ₗ[K] V') := -begin - haveI := classical.dec_eq V, - haveI := classical.dec_eq V', - let m := basis.of_vector_space K V, - let m' := basis.of_vector_space K V', - rw [←cardinal.lift_inj.1 m.mk_eq_dim, ←cardinal.lift_inj.1 m'.mk_eq_dim] at h, - rcases quotient.exact h with ⟨e⟩, - let e := (equiv.ulift.symm.trans e).trans equiv.ulift, - exact ⟨(m.repr ≪≫ₗ (finsupp.dom_lcongr e)) ≪≫ₗ m'.repr.symm⟩ -end - -/-- Two `K`-vector spaces are equivalent if their dimension is the same. -/ -def equiv_of_dim_eq_dim (h : module.rank K V₁ = module.rank K V₂) : V₁ ≃ₗ[K] V₂ := -begin - classical, - exact classical.choice (equiv_of_dim_eq_lift_dim (cardinal.lift_inj.2 h)) -end - -/-- An `n`-dimensional `K`-vector space is equivalent to `fin n → K`. -/ -def fin_dim_vectorspace_equiv (n : ℕ) - (hn : (module.rank K V) = n) : V ≃ₗ[K] (fin n → K) := -begin - have : cardinal.lift.{u} (n : cardinal.{v}) = cardinal.lift.{v} (n : cardinal.{u}), - by simp, - have hn := cardinal.lift_inj.{v u}.2 hn, - rw this at hn, - rw ←@dim_fin_fun K _ n at hn, - exact classical.choice (equiv_of_dim_eq_lift_dim hn), -end - -end module +/-! TODO: move this section to an earlier file. -/ namespace basis diff --git a/src/linear_algebra/free_algebra.lean b/src/linear_algebra/free_algebra.lean index 720c977291e64..4e805ddeda500 100644 --- a/src/linear_algebra/free_algebra.lean +++ b/src/linear_algebra/free_algebra.lean @@ -5,6 +5,7 @@ Authors: Eric Wieser -/ import linear_algebra.basis import algebra.free_algebra +import linear_algebra.dimension import linear_algebra.finsupp_vector_space /-! # Linear algebra properties of `free_algebra R X` diff --git a/src/linear_algebra/free_module/finite/basic.lean b/src/linear_algebra/free_module/finite/basic.lean index 33ea9d0df2539..c00cd6a61b431 100644 --- a/src/linear_algebra/free_module/finite/basic.lean +++ b/src/linear_algebra/free_module/finite/basic.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Riccardo Brasca -/ +import linear_algebra.dimension import linear_algebra.free_module.basic import ring_theory.finiteness diff --git a/src/linear_algebra/free_module/rank.lean b/src/linear_algebra/free_module/rank.lean index 8527af1eb7879..6a7379b9c11d2 100644 --- a/src/linear_algebra/free_module/rank.lean +++ b/src/linear_algebra/free_module/rank.lean @@ -4,8 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Riccardo Brasca -/ +import linear_algebra.dimension import linear_algebra.free_module.basic -import linear_algebra.finsupp_vector_space +import linear_algebra.invariant_basis_number /-! From 00d163e35035c3577c1c79fa53b68de17781ffc1 Mon Sep 17 00:00:00 2001 From: Alex J Best Date: Mon, 3 Apr 2023 13:33:50 +0000 Subject: [PATCH 0764/1291] feat(ring_theory/zmod): Criterion for `zmod` to be a reduced ring (#16998) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit I couldn't find a good place for this without adding some imports, so a new file seemed appropriate. Co-authored-by: Yaël Dillies --- src/algebra/squarefree.lean | 11 +++++++++++ src/data/int/basic.lean | 2 ++ src/ring_theory/zmod.lean | 27 +++++++++++++++++++++++++++ 3 files changed, 40 insertions(+) create mode 100644 src/ring_theory/zmod.lean diff --git a/src/algebra/squarefree.lean b/src/algebra/squarefree.lean index fe3338803a6b6..dd70a525e1e70 100644 --- a/src/algebra/squarefree.lean +++ b/src/algebra/squarefree.lean @@ -255,3 +255,14 @@ begin end end unique_factorization_monoid + +namespace int + +@[simp] lemma squarefree_nat_abs {n : ℤ} : squarefree n.nat_abs ↔ squarefree n := +by simp_rw [squarefree, nat_abs_surjective.forall, ←nat_abs_mul, nat_abs_dvd_iff_dvd, + is_unit_iff_nat_abs_eq, nat.is_unit_iff] + +@[simp] lemma squarefree_coe_nat {n : ℕ} : squarefree (n : ℤ) ↔ squarefree n := +by rw [←squarefree_nat_abs, nat_abs_of_nat] + +end int diff --git a/src/data/int/basic.lean b/src/data/int/basic.lean index d780c83e2a63d..f83bb630be5bd 100644 --- a/src/data/int/basic.lean +++ b/src/data/int/basic.lean @@ -186,6 +186,8 @@ variables {a b : ℤ} {n : ℕ} attribute [simp] nat_abs_of_nat nat_abs_zero nat_abs_one +lemma nat_abs_surjective : nat_abs.surjective := λ n, ⟨n, nat_abs_of_nat n⟩ + theorem nat_abs_add_le (a b : ℤ) : nat_abs (a + b) ≤ nat_abs a + nat_abs b := begin have : ∀ (a b : ℕ), nat_abs (sub_nat_nat a (nat.succ b)) ≤ nat.succ (a + b), diff --git a/src/ring_theory/zmod.lean b/src/ring_theory/zmod.lean new file mode 100644 index 0000000000000..62b6858bc9436 --- /dev/null +++ b/src/ring_theory/zmod.lean @@ -0,0 +1,27 @@ +/- +Copyright (c) 2022 Alex J. Best. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Alex J. Best +-/ +import algebra.squarefree +import data.zmod.basic +import ring_theory.int.basic + +/-! +# Ring theoretic facts about `zmod n` + +We collect a few facts about `zmod n` that need some ring theory to be proved/stated + +## Main statements + +* `is_reduced_zmod` - `zmod n` is reduced for all squarefree `n`. +-/ + +@[simp] lemma is_reduced_zmod {n : ℕ} : is_reduced (zmod n) ↔ squarefree n ∨ n = 0 := +by rw [← ring_hom.ker_is_radical_iff_reduced_of_surjective + (zmod.ring_hom_surjective $ int.cast_ring_hom $ zmod n), + zmod.ker_int_cast_ring_hom, ← is_radical_iff_span_singleton, + is_radical_iff_squarefree_or_zero, int.squarefree_coe_nat, nat.cast_eq_zero] + +instance {n : ℕ} [fact $ squarefree n] : is_reduced (zmod n) := +is_reduced_zmod.2 $ or.inl $ fact.out _ From 3cacc945118c8c637d89950af01da78307f59325 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 3 Apr 2023 16:47:49 +0000 Subject: [PATCH 0765/1291] feat(linear_algebra/dimension): free generalizations (#18722) Generalizes many results about `module.rank` from `[division_ring K]` to `[ring K] [strong_rank_condition K] [module.free K V]`. Some lemmas have been moved around in the file to make use of existing `variables` groupings. There are some lemmas about division rings that I wasn't able to weaken the assumptions on. I'll make the corresponding generalizations to `finrank` in a follow-up PR. --- src/algebra/linear_recurrence.lean | 13 ++- src/linear_algebra/dimension.lean | 147 ++++++++++++++++------------- 2 files changed, 88 insertions(+), 72 deletions(-) diff --git a/src/algebra/linear_recurrence.lean b/src/algebra/linear_recurrence.lean index be9b92321b59b..c26e070cabc6f 100644 --- a/src/algebra/linear_recurrence.lean +++ b/src/algebra/linear_recurrence.lean @@ -167,15 +167,20 @@ def tuple_succ : (fin E.order → α) →ₗ[α] (fin E.order → α) := end comm_semiring -section field +section strong_rank_condition -variables {α : Type*} [field α] (E : linear_recurrence α) +-- note: `strong_rank_condition` is the same as `nontrivial` on `comm_ring`s, but that result, +-- `comm_ring_strong_rank_condition`, is in a much later file. +variables {α : Type*} [comm_ring α] [strong_rank_condition α] (E : linear_recurrence α) /-- The dimension of `E.sol_space` is `E.order`. -/ lemma sol_space_dim : module.rank α E.sol_space = E.order := -@dim_fin_fun α _ E.order ▸ E.to_init.dim_eq +begin + letI := nontrivial_of_invariant_basis_number α, + exact @dim_fin_fun α _ _ _ E.order ▸ E.to_init.dim_eq +end -end field +end strong_rank_condition section comm_ring diff --git a/src/linear_algebra/dimension.lean b/src/linear_algebra/dimension.lean index b01c0b30f9395..b4991d5081098 100644 --- a/src/linear_algebra/dimension.lean +++ b/src/linear_algebra/dimension.lean @@ -287,6 +287,10 @@ end variables {R M} +lemma exists_mem_ne_zero_of_dim_pos {s : submodule R M} (h : 0 < module.rank R s) : + ∃ b : M, b ∈ s ∧ b ≠ 0 := +exists_mem_ne_zero_of_ne_bot $ assume eq, by rw [eq, dim_bot] at h; exact lt_irrefl _ h + /-- A linearly-independent family of vectors in a module over a non-trivial ring must be finite if the module is Noetherian. -/ lemma linear_independent.finite_of_is_noetherian [is_noetherian R M] @@ -930,31 +934,28 @@ by rw [←cardinal.lift_inj, ← (basis.singleton punit R).mk_eq_dim, cardinal.m end strong_rank_condition -section division_ring -variables [division_ring K] [add_comm_group V] [module K V] [add_comm_group V₁] [module K V₁] +section free +variables [ring K] [strong_rank_condition K] +variables [add_comm_group V] [module K V] [module.free K V] +variables [add_comm_group V'] [module K V'] [module.free K V'] +variables [add_comm_group V₁] [module K V₁] [module.free K V₁] variables {K V} -/-- If a vector space has a finite dimension, the index set of `basis.of_vector_space` is finite. -/ -lemma basis.finite_of_vector_space_index_of_dim_lt_aleph_0 (h : module.rank K V < ℵ₀) : - (basis.of_vector_space_index K V).finite := -finite_def.2 $ (basis.of_vector_space K V).nonempty_fintype_index_of_dim_lt_aleph_0 h - -variables [add_comm_group V'] [module K V'] - /-- Two vector spaces are isomorphic if they have the same dimension. -/ theorem nonempty_linear_equiv_of_lift_dim_eq (cond : cardinal.lift.{v'} (module.rank K V) = cardinal.lift.{v} (module.rank K V')) : nonempty (V ≃ₗ[K] V') := begin - let B := basis.of_vector_space K V, - let B' := basis.of_vector_space K V', + obtain ⟨⟨_, B⟩⟩ := module.free.exists_basis K V, + obtain ⟨⟨_, B'⟩⟩ := module.free.exists_basis K V', have : cardinal.lift.{v' v} (#_) = cardinal.lift.{v v'} (#_), by rw [B.mk_eq_dim'', cond, B'.mk_eq_dim''], exact (cardinal.lift_mk_eq.{v v' 0}.1 this).map (B.equiv B') end /-- Two vector spaces are isomorphic if they have the same dimension. -/ -theorem nonempty_linear_equiv_of_dim_eq (cond : module.rank K V = module.rank K V₁) : +theorem nonempty_linear_equiv_of_dim_eq + (cond : module.rank K V = module.rank K V₁) : nonempty (V ≃ₗ[K] V₁) := nonempty_linear_equiv_of_lift_dim_eq $ congr_arg _ cond @@ -985,26 +986,10 @@ theorem linear_equiv.nonempty_equiv_iff_dim_eq : nonempty (V ≃ₗ[K] V₁) ↔ module.rank K V = module.rank K V₁ := ⟨λ ⟨h⟩, linear_equiv.dim_eq h, λ h, nonempty_linear_equiv_of_dim_eq h⟩ --- TODO how far can we generalise this? --- When `s` is finite, we could prove this for any ring satisfying the strong rank condition --- using `linear_independent_le_span'` -lemma dim_span_le (s : set V) : module.rank K (span K s) ≤ #s := -begin - obtain ⟨b, hb, hsab, hlib⟩ := exists_linear_independent K s, - convert cardinal.mk_le_mk_of_subset hb, - rw [← hsab, dim_span_set hlib] -end - -lemma dim_span_of_finset (s : finset V) : - module.rank K (span K (↑s : set V)) < ℵ₀ := -calc module.rank K (span K (↑s : set V)) ≤ #(↑s : set V) : dim_span_le ↑s - ... = s.card : by rw [finset.coe_sort_coe, cardinal.mk_coe_finset] - ... < ℵ₀ : cardinal.nat_lt_aleph_0 _ - theorem dim_prod : module.rank K (V × V₁) = module.rank K V + module.rank K V₁ := begin - let b := basis.of_vector_space K V, - let c := basis.of_vector_space K V₁, + obtain ⟨⟨_, b⟩⟩ := module.free.exists_basis K V, + obtain ⟨⟨_, c⟩⟩ := module.free.exists_basis K V₁, rw [← cardinal.lift_inj, ← (basis.prod b c).mk_eq_dim, cardinal.lift_add, ← cardinal.mk_ulift, @@ -1016,38 +1001,93 @@ begin end section fintype -variables [∀i, add_comm_group (φ i)] [∀i, module K (φ i)] +variables [∀i, add_comm_group (φ i)] [∀i, module K (φ i)] [∀i, module.free K (φ i)] open linear_map -lemma dim_pi [finite η] : module.rank K (Πi, φ i) = cardinal.sum (λi, module.rank K (φ i)) := +lemma dim_pi [nontrivial K] [finite η] : + module.rank K (Πi, φ i) = cardinal.sum (λi, module.rank K (φ i)) := begin casesI nonempty_fintype η, - let b := assume i, basis.of_vector_space K (φ i), + let b := λ i, (module.free.exists_basis K (φ i)).some.2, let this : basis (Σ j, _) K (Π j, φ j) := pi.basis b, rw [← cardinal.lift_inj, ← this.mk_eq_dim], - simp [← (b _).mk_range_eq_dim] + simp_rw [cardinal.mk_sigma, cardinal.lift_sum, ←(b _).mk_range_eq_dim, + cardinal.mk_range_eq _ (b _).injective], end variable [fintype η] -lemma dim_fun {V η : Type u} [fintype η] [add_comm_group V] [module K V] : +lemma dim_fun {V η : Type u} [nontrivial K] [fintype η] [add_comm_group V] [module K V] + [module.free K V] : module.rank K (η → V) = fintype.card η * module.rank K V := by rw [dim_pi, cardinal.sum_const', cardinal.mk_fintype] -lemma dim_fun_eq_lift_mul : +lemma dim_fun_eq_lift_mul [nontrivial K] : module.rank K (η → V) = (fintype.card η : cardinal.{max u₁' v}) * cardinal.lift.{u₁'} (module.rank K V) := by rw [dim_pi, cardinal.sum_const, cardinal.mk_fintype, cardinal.lift_nat_cast] -lemma dim_fun' : module.rank K (η → K) = fintype.card η := +lemma dim_fun' [nontrivial K] : module.rank K (η → K) = fintype.card η := by rw [dim_fun_eq_lift_mul, dim_self, cardinal.lift_one, mul_one, cardinal.nat_cast_inj] -lemma dim_fin_fun (n : ℕ) : module.rank K (fin n → K) = n := +lemma dim_fin_fun [nontrivial K] (n : ℕ) : module.rank K (fin n → K) = n := by simp [dim_fun'] end fintype +lemma finsupp.dim_eq {ι : Type v} : module.rank K (ι →₀ V) = #ι * module.rank K V := +begin + obtain ⟨⟨_, bs⟩⟩ := module.free.exists_basis K V, + rw [← bs.mk_eq_dim'', ← (finsupp.basis (λa:ι, bs)).mk_eq_dim'', + cardinal.mk_sigma, cardinal.sum_const'] +end + +-- TODO: merge with the `finrank` content +/-- An `n`-dimensional `K`-vector space is equivalent to `fin n → K`. -/ +def fin_dim_vectorspace_equiv (n : ℕ) + (hn : (module.rank K V) = n) : V ≃ₗ[K] (fin n → K) := +begin + haveI := nontrivial_of_invariant_basis_number K, + have : cardinal.lift.{u} (n : cardinal.{v}) = cardinal.lift.{v} (n : cardinal.{u}), + by simp, + have hn := cardinal.lift_inj.{v u}.2 hn, + rw this at hn, + rw ←@dim_fin_fun K _ _ _ n at hn, + haveI : module.free K (fin n → K) := module.free.pi _ _, + exact classical.choice (nonempty_linear_equiv_of_lift_dim_eq hn), +end + +end free + +section division_ring +variables [division_ring K] +variables [add_comm_group V] [module K V] +variables [add_comm_group V'] [module K V'] +variables [add_comm_group V₁] [module K V₁] +variables {K V} + +/-- If a vector space has a finite dimension, the index set of `basis.of_vector_space` is finite. -/ +lemma basis.finite_of_vector_space_index_of_dim_lt_aleph_0 (h : module.rank K V < ℵ₀) : + (basis.of_vector_space_index K V).finite := +finite_def.2 $ (basis.of_vector_space K V).nonempty_fintype_index_of_dim_lt_aleph_0 h + +-- TODO how far can we generalise this? +-- When `s` is finite, we could prove this for any ring satisfying the strong rank condition +-- using `linear_independent_le_span'` +lemma dim_span_le (s : set V) : module.rank K (span K s) ≤ #s := +begin + obtain ⟨b, hb, hsab, hlib⟩ := exists_linear_independent K s, + convert cardinal.mk_le_mk_of_subset hb, + rw [← hsab, dim_span_set hlib] +end + +lemma dim_span_of_finset (s : finset V) : + module.rank K (span K (↑s : set V)) < ℵ₀ := +calc module.rank K (span K (↑s : set V)) ≤ #(↑s : set V) : dim_span_le ↑s + ... = s.card : by rw [finset.coe_sort_coe, cardinal.mk_coe_finset] + ... < ℵ₀ : cardinal.nat_lt_aleph_0 _ + theorem dim_quotient_add_dim (p : submodule K V) : module.rank K (V ⧸ p) + module.rank K p = module.rank K V := by classical; exact let ⟨f⟩ := quotient_prod_linear_equiv p in dim_prod.symm.trans f.dim_eq @@ -1126,10 +1166,6 @@ by { rw [← dim_sup_add_dim_inf_eq], exact self_le_add_right _ _ } end -lemma exists_mem_ne_zero_of_dim_pos {s : submodule K V} (h : 0 < module.rank K s) : - ∃ b : V, b ∈ s ∧ b ≠ 0 := -exists_mem_ne_zero_of_ne_bot $ assume eq, by rw [eq, dim_bot] at h; exact lt_irrefl _ h - end division_ring section rank @@ -1250,7 +1286,7 @@ begin have h : (K ∙ v₀) = ⊤, { ext, simp [mem_span_singleton, hv₀] }, rw [←dim_top, ←h], - convert dim_span_le _, + refine (dim_span_le _).trans_eq _, simp } end @@ -1355,28 +1391,3 @@ end end division_ring end module - -section field -variables [field K] [add_comm_group V] [module K V] - -lemma finsupp.dim_eq {ι : Type v} : module.rank K (ι →₀ V) = #ι * module.rank K V := -begin - let bs := basis.of_vector_space K V, - rw [← bs.mk_eq_dim'', ← (finsupp.basis (λa:ι, bs)).mk_eq_dim'', - cardinal.mk_sigma, cardinal.sum_const'] -end - --- TODO: merge with the `finrank` content -/-- An `n`-dimensional `K`-vector space is equivalent to `fin n → K`. -/ -def fin_dim_vectorspace_equiv (n : ℕ) - (hn : (module.rank K V) = n) : V ≃ₗ[K] (fin n → K) := -begin - have : cardinal.lift.{u} (n : cardinal.{v}) = cardinal.lift.{v} (n : cardinal.{u}), - by simp, - have hn := cardinal.lift_inj.{v u}.2 hn, - rw this at hn, - rw ←@dim_fin_fun K _ n at hn, - exact classical.choice (nonempty_linear_equiv_of_lift_dim_eq hn), -end - -end field From e655e4ea5c6d02854696f97494997ba4c31be802 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 3 Apr 2023 20:45:06 +0000 Subject: [PATCH 0766/1291] feat(algebra/group_power/lemmas): Induction principle for powers (#18668) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit A property holds for all powers of `g` if it is holds for `1` and is preserved under multiplication and division by `g`. Also rename `subgroup.zpowers_subset`/`add_subgroup.zmultiples_subset` to `subgroup.zpowers_le_of_mem`/`add_subgroup.zmultiples_le_of_mem` because there is no `⊆` in the statement. The motivation is the Cauchy-Davenport theorem: https://github.com/leanprover-community/mathlib/blob/321b67021163ac504c6cfa35d5678a47b357869d/src/combinatorics/additive/cauchy_davenport.lean#L176-L181 --- src/algebra/group_power/lemmas.lean | 34 ++++++++++++++++++++++ src/group_theory/subgroup/pointwise.lean | 16 ++++++---- src/group_theory/subgroup/zpowers.lean | 16 +++++++--- src/group_theory/submonoid/membership.lean | 3 ++ src/ring_theory/int/basic.lean | 4 +-- src/ring_theory/roots_of_unity.lean | 2 +- 6 files changed, 62 insertions(+), 13 deletions(-) diff --git a/src/algebra/group_power/lemmas.lean b/src/algebra/group_power/lemmas.lean index 89259ef81f0ac..aad2d21b28729 100644 --- a/src/algebra/group_power/lemmas.lean +++ b/src/algebra/group_power/lemmas.lean @@ -185,6 +185,40 @@ by rw [zpow_add, zpow_one] theorem zpow_bit1 (a : G) (n : ℤ) : a ^ bit1 n = a ^ n * a ^ n * a := by rw [bit1, zpow_add, zpow_bit0, zpow_one] +/-- To show a property of all powers of `g` it suffices to show it is closed under multiplication +by `g` and `g⁻¹` on the left. For subgroups generated by more than one element, see +`subgroup.closure_induction_left`. -/ +@[to_additive "To show a property of all multiples of `g` it suffices to show it is closed under +addition by `g` and `-g` on the left. For additive subgroups generated by more than one element, see +`add_subgroup.closure_induction_left`."] +lemma zpow_induction_left {g : G} {P : G → Prop} (h_one : P (1 : G)) + (h_mul : ∀ a, P a → P (g * a)) (h_inv : ∀ a, P a → P (g⁻¹ * a)) (n : ℤ) : P (g ^ n) := +begin + induction n using int.induction_on with n ih n ih, + { rwa zpow_zero }, + { rw [add_comm, zpow_add, zpow_one], + exact h_mul _ ih }, + { rw [sub_eq_add_neg, add_comm, zpow_add, zpow_neg_one], + exact h_inv _ ih } +end + +/-- To show a property of all powers of `g` it suffices to show it is closed under multiplication +by `g` and `g⁻¹` on the right. For subgroups generated by more than one element, see +`subgroup.closure_induction_right`. -/ +@[to_additive "To show a property of all multiples of `g` it suffices to show it is closed under +addition by `g` and `-g` on the right. For additive subgroups generated by more than one element, +see `add_subgroup.closure_induction_right`."] +lemma zpow_induction_right {g : G} {P : G → Prop} (h_one : P (1 : G)) + (h_mul : ∀ a, P a → P (a * g)) (h_inv : ∀ a, P a → P (a * g⁻¹)) (n : ℤ) : P (g ^ n) := +begin + induction n using int.induction_on with n ih n ih, + { rwa zpow_zero }, + { rw zpow_add_one, + exact h_mul _ ih }, + { rw zpow_sub_one, + exact h_inv _ ih } +end + end group /-! diff --git a/src/group_theory/subgroup/pointwise.lean b/src/group_theory/subgroup/pointwise.lean index 3b596c7dd8083..8b31a7d37c61d 100644 --- a/src/group_theory/subgroup/pointwise.lean +++ b/src/group_theory/subgroup/pointwise.lean @@ -55,15 +55,19 @@ begin { simp only [true_and, coe_to_submonoid, union_subset_iff, subset_closure, inv_subset_closure] } end -@[to_additive] lemma closure_induction_left {p : G → Prop} {x : G} - (h : x ∈ closure s) (H1 : p 1) (Hmul : ∀ (x ∈ s) y, p y → p (x * y)) - (Hinv : ∀ (x ∈ s) y, p y → p (x⁻¹ * y)) : p x := +/-- For subgroups generated by a single element, see the simpler `zpow_induction_left`. -/ +@[to_additive "For additive subgroups generated by a single element, see the simpler +`zsmul_induction_left`."] +lemma closure_induction_left {p : G → Prop} {x : G} (h : x ∈ closure s) (H1 : p 1) + (Hmul : ∀ (x ∈ s) y, p y → p (x * y)) (Hinv : ∀ (x ∈ s) y, p y → p (x⁻¹ * y)) : p x := let key := (closure_to_submonoid s).le in submonoid.closure_induction_left (key h) H1 $ λ x hx, hx.elim (Hmul x) $ λ hx y hy, (congr_arg _ $ inv_inv x).mp $ Hinv x⁻¹ hx y hy -@[to_additive] lemma closure_induction_right {p : G → Prop} {x : G} - (h : x ∈ closure s) (H1 : p 1) (Hmul : ∀ x (y ∈ s), p x → p (x * y)) - (Hinv : ∀ x (y ∈ s), p x → p (x * y⁻¹)) : p x := +/-- For subgroups generated by a single element, see the simpler `zpow_induction_right`. -/ +@[to_additive "For additive subgroups generated by a single element, see the simpler +`zsmul_induction_right`."] +lemma closure_induction_right {p : G → Prop} {x : G} (h : x ∈ closure s) (H1 : p 1) + (Hmul : ∀ x (y ∈ s), p x → p (x * y)) (Hinv : ∀ x (y ∈ s), p x → p (x * y⁻¹)) : p x := let key := (closure_to_submonoid s).le in submonoid.closure_induction_right (key h) H1 $ λ x y hy, hy.elim (Hmul x y) $ λ hy hx, (congr_arg _ $ inv_inv y).mp $ Hinv x y⁻¹ hy hx diff --git a/src/group_theory/subgroup/zpowers.lean b/src/group_theory/subgroup/zpowers.lean index ca2ee8f68e936..6de48373ae8b7 100644 --- a/src/group_theory/subgroup/zpowers.lean +++ b/src/group_theory/subgroup/zpowers.lean @@ -29,14 +29,13 @@ subgroup.copy (zpowers_hom G g).range (set.range ((^) g : ℤ → G)) rfl @[simp] lemma mem_zpowers (g : G) : g ∈ zpowers g := ⟨1, zpow_one _⟩ +@[norm_cast] lemma coe_zpowers (g : G) : ↑(zpowers g) = set.range (λ n : ℤ, g ^ n) := rfl + lemma zpowers_eq_closure (g : G) : zpowers g = closure {g} := by { ext, exact mem_closure_singleton.symm } @[simp] lemma range_zpowers_hom (g : G) : (zpowers_hom G g).range = zpowers g := rfl -lemma zpowers_subset {a : G} {K : subgroup G} (h : a ∈ K) : zpowers a ≤ K := -λ x hx, match x, hx with _, ⟨i, rfl⟩ := K.zpow_mem h i end - lemma mem_zpowers_iff {g h : G} : h ∈ zpowers g ↔ ∃ (k : ℤ), g ^ k = h := iff.rfl @@ -78,9 +77,9 @@ add_subgroup.copy (zmultiples_hom A a).range (set.range ((• a) : ℤ → A)) r attribute [to_additive add_subgroup.zmultiples] subgroup.zpowers attribute [to_additive add_subgroup.mem_zmultiples] subgroup.mem_zpowers +attribute [to_additive add_subgroup.coe_zmultiples] subgroup.coe_zpowers attribute [to_additive add_subgroup.zmultiples_eq_closure] subgroup.zpowers_eq_closure attribute [to_additive add_subgroup.range_zmultiples_hom] subgroup.range_zpowers_hom -attribute [to_additive add_subgroup.zmultiples_subset] subgroup.zpowers_subset attribute [to_additive add_subgroup.mem_zmultiples_iff] subgroup.mem_zpowers_iff attribute [to_additive add_subgroup.zsmul_mem_zmultiples] subgroup.zpow_mem_zpowers attribute [to_additive add_subgroup.nsmul_mem_zmultiples] subgroup.npow_mem_zpowers @@ -140,6 +139,7 @@ begin end namespace subgroup +variables {s : set G} {g : G} @[to_additive zmultiples_is_commutative] instance zpowers_is_commutative (g : G) : (zpowers g).is_commutative := @@ -150,9 +150,17 @@ instance zpowers_is_commutative (g : G) : (zpowers g).is_commutative := lemma zpowers_le {g : G} {H : subgroup G} : zpowers g ≤ H ↔ g ∈ H := by rw [zpowers_eq_closure, closure_le, set.singleton_subset_iff, set_like.mem_coe] +alias zpowers_le ↔ _ zpowers_le_of_mem +alias add_subgroup.zmultiples_le ↔ _ _root_.add_subgroup.zmultiples_le_of_mem + +attribute [to_additive zmultiples_le_of_mem] zpowers_le_of_mem + @[simp, to_additive zmultiples_eq_bot] lemma zpowers_eq_bot {g : G} : zpowers g = ⊥ ↔ g = 1 := by rw [eq_bot_iff, zpowers_le, mem_bot] +@[to_additive zmultiples_ne_bot] lemma zpowers_ne_bot : zpowers g ≠ ⊥ ↔ g ≠ 1 := +zpowers_eq_bot.not + @[simp, to_additive zmultiples_zero_eq_bot] lemma zpowers_one_eq_bot : subgroup.zpowers (1 : G) = ⊥ := subgroup.zpowers_eq_bot.mpr rfl diff --git a/src/group_theory/submonoid/membership.lean b/src/group_theory/submonoid/membership.lean index f275719bbacb0..d25b78b2e4ce6 100644 --- a/src/group_theory/submonoid/membership.lean +++ b/src/group_theory/submonoid/membership.lean @@ -346,6 +346,8 @@ set.ext (λ n, exists_congr $ λ i, by simp; refl) @[simp] lemma mem_powers (n : M) : n ∈ powers n := ⟨1, pow_one _⟩ +@[norm_cast] lemma coe_powers (x : M) : ↑(powers x) = set.range (λ n : ℕ, x ^ n) := rfl + lemma mem_powers_iff (x z : M) : x ∈ powers z ↔ ∃ n : ℕ, z ^ n = x := iff.rfl lemma powers_eq_closure (n : M) : powers n = closure {n} := @@ -479,6 +481,7 @@ set.ext (λ n, exists_congr $ λ i, by simp; refl) attribute [to_additive multiples] submonoid.powers attribute [to_additive mem_multiples] submonoid.mem_powers +attribute [to_additive coe_multiples] submonoid.coe_powers attribute [to_additive mem_multiples_iff] submonoid.mem_powers_iff attribute [to_additive multiples_eq_closure] submonoid.powers_eq_closure attribute [to_additive multiples_subset] submonoid.powers_subset diff --git a/src/ring_theory/int/basic.lean b/src/ring_theory/int/basic.lean index 086a97a34cc12..f03522efeb169 100644 --- a/src/ring_theory/int/basic.lean +++ b/src/ring_theory/int/basic.lean @@ -358,8 +358,8 @@ namespace int lemma zmultiples_nat_abs (a : ℤ) : add_subgroup.zmultiples (a.nat_abs : ℤ) = add_subgroup.zmultiples a := le_antisymm - (add_subgroup.zmultiples_subset (mem_zmultiples_iff.mpr (dvd_nat_abs.mpr (dvd_refl a)))) - (add_subgroup.zmultiples_subset (mem_zmultiples_iff.mpr (nat_abs_dvd.mpr (dvd_refl a)))) + (add_subgroup.zmultiples_le_of_mem (mem_zmultiples_iff.mpr (dvd_nat_abs.mpr (dvd_refl a)))) + (add_subgroup.zmultiples_le_of_mem (mem_zmultiples_iff.mpr (nat_abs_dvd.mpr (dvd_refl a)))) lemma span_nat_abs (a : ℤ) : ideal.span ({a.nat_abs} : set ℤ) = ideal.span {a} := by { rw ideal.span_singleton_eq_span_singleton, exact (associated_nat_abs _).symm } diff --git a/src/ring_theory/roots_of_unity.lean b/src/ring_theory/roots_of_unity.lean index 938d0c449973c..d5fe89ae3a8ca 100644 --- a/src/ring_theory/roots_of_unity.lean +++ b/src/ring_theory/roots_of_unity.lean @@ -679,7 +679,7 @@ begin haveI F : fintype (subgroup.zpowers ζ) := fintype.of_equiv _ (h.zmod_equiv_zpowers).to_equiv, refine @set.eq_of_subset_of_card_le Rˣ (subgroup.zpowers ζ) (roots_of_unity k R) F (roots_of_unity.fintype R k) - (subgroup.zpowers_subset $ show ζ ∈ roots_of_unity k R, from h.pow_eq_one) _, + (subgroup.zpowers_le_of_mem $ show ζ ∈ roots_of_unity k R, from h.pow_eq_one) _, calc fintype.card (roots_of_unity k R) ≤ k : card_roots_of_unity R k ... = fintype.card (zmod k) : (zmod.card k).symm From 08a4542bec7242a5c60f179e4e49de8c0d677b1b Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Tue, 4 Apr 2023 06:24:46 +0000 Subject: [PATCH 0767/1291] feat(measure_theory/function): Functions locally integrable on a set (#18673) This PR adds a definition for local integrability of a function on a set (a relative version of the existing "locally integrable"), and proves some elementary properties of this definition (in particular, functions continuous on `s` are locally integrable on `s`). --- .../function/locally_integrable.lean | 360 +++++++++++++----- .../integral/integrable_on.lean | 15 +- .../integral/interval_integral.lean | 5 +- src/measure_theory/integral/layercake.lean | 2 +- src/measure_theory/integral/set_integral.lean | 4 +- 5 files changed, 285 insertions(+), 101 deletions(-) diff --git a/src/measure_theory/function/locally_integrable.lean b/src/measure_theory/function/locally_integrable.lean index 54eff2ef7aa62..af86545e34b22 100644 --- a/src/measure_theory/function/locally_integrable.lean +++ b/src/measure_theory/function/locally_integrable.lean @@ -9,15 +9,17 @@ import measure_theory.integral.integrable_on # Locally integrable functions A function is called *locally integrable* (`measure_theory.locally_integrable`) if it is integrable -on a neighborhood of every point. +on a neighborhood of every point. More generally, it is *locally integrable on `s`* if it is +locally integrable on a neighbourhood within `s` of any point of `s`. -This file contains properties of locally integrable functions and integrability results +This file contains properties of locally integrable functions, and integrability results on compact sets. ## Main statements * `continuous.locally_integrable`: A continuous function is locally integrable. - +* `continuous_on.locally_integrable_on`: A function which is continuous on `s` is locally + integrable on `s`. -/ open measure_theory measure_theory.measure set function topological_space @@ -25,19 +27,141 @@ open_locale topology interval variables {X Y E R : Type*} [measurable_space X] [topological_space X] variables [measurable_space Y] [topological_space Y] -variables [normed_add_comm_group E] {f : X → E} {μ : measure X} +variables [normed_add_comm_group E] {f : X → E} {μ : measure X} {s : set X} namespace measure_theory -/-- A function `f : X → E` is locally integrable if it is integrable on a neighborhood of every +section locally_integrable_on + +/-- A function `f : X → E` is *locally integrable on s*, for `s ⊆ X`, if for every `x ∈ s` there is +a neighbourhood of `x` within `s` on which `f` is integrable. (Note this is, in general, strictly +weaker than local integrability with respect to `μ.restrict s`.) -/ +def locally_integrable_on (f : X → E) (s : set X) (μ : measure X . volume_tac) : Prop := +∀ (x : X), x ∈ s → integrable_at_filter f (𝓝[s] x) μ + +lemma locally_integrable_on.mono + (hf : measure_theory.locally_integrable_on f s μ) {t : set X} (hst : t ⊆ s) : + locally_integrable_on f t μ := +λ x hx, (hf x $ hst hx).filter_mono (nhds_within_mono x hst) + +lemma locally_integrable_on.norm (hf : locally_integrable_on f s μ) : + locally_integrable_on (λ x, ‖f x‖) s μ := +λ t ht, let ⟨U, hU_nhd, hU_int⟩ := hf t ht in ⟨U, hU_nhd, hU_int.norm⟩ + +lemma integrable_on.locally_integrable_on (hf : integrable_on f s μ) : + locally_integrable_on f s μ := +λ x hx, ⟨s, self_mem_nhds_within, hf⟩ + +/-- If a function is locally integrable on a compact set, then it is integrable on that set. -/ +lemma locally_integrable_on.integrable_on_is_compact + (hf : locally_integrable_on f s μ) (hs : is_compact s) : + integrable_on f s μ := +is_compact.induction_on hs integrable_on_empty (λ u v huv hv, hv.mono_set huv) + (λ u v hu hv, integrable_on_union.mpr ⟨hu, hv⟩) hf + +lemma locally_integrable_on.integrable_on_compact_subset + (hf : locally_integrable_on f s μ) {t : set X} (hst : t ⊆ s) (ht : is_compact t) : + integrable_on f t μ := +(hf.mono hst).integrable_on_is_compact ht + +lemma locally_integrable_on.ae_strongly_measurable [second_countable_topology X] + (hf : locally_integrable_on f s μ) : + ae_strongly_measurable f (μ.restrict s) := +begin + have : ∀ (x : s), ∃ u, is_open u ∧ x.1 ∈ u ∧ integrable_on f (u ∩ s) μ, + { rintro ⟨x, hx⟩, + rcases hf x hx with ⟨t, ht, h't⟩, + rcases mem_nhds_within.1 ht with ⟨u, u_open, x_mem, u_sub⟩, + refine ⟨u, u_open, x_mem, h't.mono_set u_sub⟩ }, + choose u u_open xu hu using this, + obtain ⟨T, T_count, hT⟩ : ∃ (T : set s), T.countable ∧ s = (⋃ (i : T), u i ∩ s), + { have : s ⊆ (⋃ (x : s), u x), from λ y hy, mem_Union_of_mem ⟨y, hy⟩ (xu ⟨y, hy⟩), + obtain ⟨T, hT_count, hT_un⟩ := is_open_Union_countable u u_open, + refine ⟨T, hT_count, _⟩, + rw [←hT_un, bUnion_eq_Union] at this, + rw [←Union_inter, eq_comm, inter_eq_right_iff_subset], + exact this }, + haveI : countable T, from countable_coe_iff.mpr T_count, + rw [hT, ae_strongly_measurable_Union_iff], + exact λ (i : T), (hu i).ae_strongly_measurable, +end + +/-- If `s` is either open, or closed, then `f` is locally integrable on `s` iff it is integrable on +every compact subset contained in `s`. -/ +lemma locally_integrable_on_iff [locally_compact_space X] [t2_space X] + (hs : is_closed s ∨ is_open s) : + locally_integrable_on f s μ ↔ ∀ (k : set X) (hk : k ⊆ s), is_compact k → integrable_on f k μ := +begin + -- The correct condition is that `s` be *locally closed*, i.e. for every `x ∈ s` there is some + -- `U ∈ 𝓝 x` such that `U ∩ s` is closed. But mathlib doesn't have locally closed sets yet. + refine ⟨λ hf k hk, hf.integrable_on_compact_subset hk, λ hf x hx, _⟩, + cases hs, + { exact let ⟨K, hK, h2K⟩ := exists_compact_mem_nhds x in ⟨_, inter_mem_nhds_within s h2K, + hf _ (inter_subset_left _ _) (is_compact_of_is_closed_subset hK (hs.inter hK.is_closed) + (inter_subset_right _ _))⟩ }, + { obtain ⟨K, hK, h2K, h3K⟩ := exists_compact_subset hs hx, + refine ⟨K, _, hf K h3K hK⟩, + simpa only [is_open.nhds_within_eq hs hx, interior_eq_nhds'] using h2K } +end + +end locally_integrable_on + +/-- A function `f : X → E` is *locally integrable* if it is integrable on a neighborhood of every point. In particular, it is integrable on all compact sets, see `locally_integrable.integrable_on_is_compact`. -/ def locally_integrable (f : X → E) (μ : measure X . volume_tac) : Prop := ∀ (x : X), integrable_at_filter f (𝓝 x) μ +lemma locally_integrable_on_univ : locally_integrable_on f univ μ ↔ locally_integrable f μ := +by simpa only [locally_integrable_on, nhds_within_univ, mem_univ, true_implies_iff] + +lemma locally_integrable.locally_integrable_on (hf : locally_integrable f μ) (s : set X) : + locally_integrable_on f s μ := +λ x hx, (hf x).filter_mono nhds_within_le_nhds + lemma integrable.locally_integrable (hf : integrable f μ) : locally_integrable f μ := λ x, hf.integrable_at_filter _ +/-- If `f` is locally integrable with respect to `μ.restrict s`, it is locally integrable on `s`. +(See `locally_integrable_on_iff_locally_integrable_restrict` for an iff statement when `s` is +closed.) -/ +lemma locally_integrable_on_of_locally_integrable_restrict [opens_measurable_space X] + (hf : locally_integrable f (μ.restrict s)) : + locally_integrable_on f s μ := +begin + intros x hx, + obtain ⟨t, ht_mem, ht_int⟩ := hf x, + obtain ⟨u, hu_sub, hu_o, hu_mem⟩ := mem_nhds_iff.mp ht_mem, + refine ⟨_, inter_mem_nhds_within s (hu_o.mem_nhds hu_mem), _⟩, + simpa only [integrable_on, measure.restrict_restrict hu_o.measurable_set, inter_comm] + using ht_int.mono_set hu_sub, +end + +/-- If `s` is closed, being locally integrable on `s` wrt `μ` is equivalent to being locally +integrable with respect to `μ.restrict s`. For the one-way implication without assuming `s` closed, +see `locally_integrable_on_of_locally_integrable_restrict`. -/ +lemma locally_integrable_on_iff_locally_integrable_restrict [opens_measurable_space X] + (hs : is_closed s) : + locally_integrable_on f s μ ↔ locally_integrable f (μ.restrict s) := +begin + refine ⟨λ hf x, _, locally_integrable_on_of_locally_integrable_restrict⟩, + by_cases h : x ∈ s, + { obtain ⟨t, ht_nhds, ht_int⟩ := hf x h, + obtain ⟨u, hu_o, hu_x, hu_sub⟩ := mem_nhds_within.mp ht_nhds, + refine ⟨u, hu_o.mem_nhds hu_x, _⟩, + rw [integrable_on, restrict_restrict hu_o.measurable_set], + exact ht_int.mono_set hu_sub }, + { rw ←is_open_compl_iff at hs, + refine ⟨sᶜ, hs.mem_nhds h, _⟩, + rw [integrable_on, restrict_restrict, inter_comm, inter_compl_self, ←integrable_on], + exacts [integrable_on_empty, hs.measurable_set] }, +end + +/-- If a function is locally integrable, then it is integrable on any compact set. -/ +lemma locally_integrable.integrable_on_is_compact {k : set X} (hf : locally_integrable f μ) + (hk : is_compact k) : integrable_on f k μ := +(hf.locally_integrable_on k).integrable_on_is_compact hk + /-- If a function is locally integrable, then it is integrable on an open neighborhood of any compact set. -/ lemma locally_integrable.integrable_on_nhds_is_compact (hf : locally_integrable f μ) {k : set X} @@ -55,40 +179,14 @@ begin exact ⟨v, nhds_within_le_nhds (v_open.mem_nhds xv), v, v_open, subset.rfl, hu.mono_set vu⟩ } end -/-- If a function is locally integrable, then it is integrable on any compact set. -/ -lemma locally_integrable.integrable_on_is_compact {k : set X} (hf : locally_integrable f μ) - (hk : is_compact k) : integrable_on f k μ := -begin - rcases hf.integrable_on_nhds_is_compact hk with ⟨u, u_open, ku, hu⟩, - exact hu.mono_set ku -end - lemma locally_integrable_iff [locally_compact_space X] : locally_integrable f μ ↔ ∀ (k : set X), is_compact k → integrable_on f k μ := -begin - refine ⟨λ hf k hk, hf.integrable_on_is_compact hk, λ hf x, _⟩, - obtain ⟨K, hK, h2K⟩ := exists_compact_mem_nhds x, - exact ⟨K, h2K, hf K hK⟩, -end +⟨λ hf k hk, hf.integrable_on_is_compact hk, + λ hf x, let ⟨K, hK, h2K⟩ := exists_compact_mem_nhds x in ⟨K, h2K, hf K hK⟩⟩ lemma locally_integrable.ae_strongly_measurable [second_countable_topology X] (hf : locally_integrable f μ) : ae_strongly_measurable f μ := -begin - have : ∀ x, ∃ u, is_open u ∧ x ∈ u ∧ integrable_on f u μ, - { assume x, - rcases hf x with ⟨s, hs, h's⟩, - rcases mem_nhds_iff.1 hs with ⟨u, us, u_open, xu⟩, - exact ⟨u, u_open, xu, h's.mono_set us⟩ }, - choose u u_open xu hu using this, - obtain ⟨T, T_count, hT⟩ : ∃ (T : set X), T.countable ∧ (⋃ (i : T), u i) = univ, - { have : (⋃ x, u x) = univ, from eq_univ_of_forall (λ x, mem_Union_of_mem x (xu x)), - rw ← this, - simp only [Union_coe_set, subtype.coe_mk], - exact is_open_Union_countable u u_open }, - haveI : countable T, from countable_coe_iff.mpr T_count, - rw [← @restrict_univ _ _ μ, ← hT, ae_strongly_measurable_Union_iff], - exact λ i, (hu i).ae_strongly_measurable, -end +by simpa only [restrict_univ] using (locally_integrable_on_univ.mpr hf).ae_strongly_measurable lemma locally_integrable_const [is_locally_finite_measure μ] (c : E) : locally_integrable (λ x, c) μ := @@ -99,6 +197,10 @@ begin simp only [h'U, integrable_on_const, or_true], end +lemma locally_integrable_on_const [is_locally_finite_measure μ] (c : E) : + locally_integrable_on (λ x, c) s μ := +(locally_integrable_const c).locally_integrable_on s + lemma locally_integrable.indicator (hf : locally_integrable f μ) {s : set X} (hs : measurable_set s) : locally_integrable (s.indicator f) μ := begin @@ -124,63 +226,9 @@ begin simp only [mem_preimage, homeomorph.symm_apply_apply] } end -section mul - -variables [opens_measurable_space X] [normed_ring R] [second_countable_topology_either X R] - {A K : set X} {g g' : X → R} - -lemma integrable_on.mul_continuous_on_of_subset - (hg : integrable_on g A μ) (hg' : continuous_on g' K) - (hA : measurable_set A) (hK : is_compact K) (hAK : A ⊆ K) : - integrable_on (λ x, g x * g' x) A μ := -begin - rcases is_compact.exists_bound_of_continuous_on hK hg' with ⟨C, hC⟩, - rw [integrable_on, ← mem_ℒp_one_iff_integrable] at hg ⊢, - have : ∀ᵐ x ∂(μ.restrict A), ‖g x * g' x‖ ≤ C * ‖g x‖, - { filter_upwards [ae_restrict_mem hA] with x hx, - refine (norm_mul_le _ _).trans _, - rw mul_comm, - apply mul_le_mul_of_nonneg_right (hC x (hAK hx)) (norm_nonneg _), }, - exact mem_ℒp.of_le_mul hg (hg.ae_strongly_measurable.mul $ - (hg'.mono hAK).ae_strongly_measurable hA) this, -end - -lemma integrable_on.mul_continuous_on [t2_space X] - (hg : integrable_on g K μ) (hg' : continuous_on g' K) (hK : is_compact K) : - integrable_on (λ x, g x * g' x) K μ := -hg.mul_continuous_on_of_subset hg' hK.measurable_set hK (subset.refl _) - -lemma integrable_on.continuous_on_mul_of_subset - (hg : continuous_on g K) (hg' : integrable_on g' A μ) - (hK : is_compact K) (hA : measurable_set A) (hAK : A ⊆ K) : - integrable_on (λ x, g x * g' x) A μ := -begin - rcases is_compact.exists_bound_of_continuous_on hK hg with ⟨C, hC⟩, - rw [integrable_on, ← mem_ℒp_one_iff_integrable] at hg' ⊢, - have : ∀ᵐ x ∂(μ.restrict A), ‖g x * g' x‖ ≤ C * ‖g' x‖, - { filter_upwards [ae_restrict_mem hA] with x hx, - refine (norm_mul_le _ _).trans _, - apply mul_le_mul_of_nonneg_right (hC x (hAK hx)) (norm_nonneg _), }, - exact mem_ℒp.of_le_mul hg' (((hg.mono hAK).ae_strongly_measurable hA).mul - hg'.ae_strongly_measurable) this, -end - -lemma integrable_on.continuous_on_mul [t2_space X] - (hg : continuous_on g K) (hg' : integrable_on g' K μ) (hK : is_compact K) : - integrable_on (λ x, g x * g' x) K μ := -hg'.continuous_on_mul_of_subset hg hK hK.measurable_set subset.rfl - -end mul - end measure_theory -open measure_theory -/-- If a function is integrable at `𝓝[s] x` for each point `x` of a compact set `s`, then it is -integrable on `s`. -/ -lemma is_compact.integrable_on_of_nhds_within {K : set X} (hK : is_compact K) - (hf : ∀ x ∈ K, integrable_at_filter f (𝓝[K] x) μ) : integrable_on f K μ := -is_compact.induction_on hK integrable_on_empty (λ s t hst ht, ht.mono_set hst) - (λ s t hs ht, hs.union ht) hf +open measure_theory section borel @@ -192,6 +240,13 @@ lemma continuous.locally_integrable [second_countable_topology_either X E] (hf : continuous f) : locally_integrable f μ := hf.integrable_at_nhds +/-- A function `f` continuous on a set `K` is locally integrable on this set with respect +to any locally finite measure. -/ +lemma continuous_on.locally_integrable_on [second_countable_topology_either X E] + (hf : continuous_on f K) (hK : measurable_set K) : + locally_integrable_on f K μ := +λ x hx, hf.integrable_at_nhds_within hK hx + variables [metrizable_space X] /-- A function `f` continuous on a compact set `K` is integrable on this set with respect to any @@ -200,7 +255,7 @@ lemma continuous_on.integrable_on_compact (hK : is_compact K) (hf : continuous_o integrable_on f K μ := begin letI := metrizable_space_metric X, - apply hK.integrable_on_of_nhds_within (λ x hx, _), + refine locally_integrable_on.integrable_on_is_compact (λ x hx, _) hK, exact hf.integrable_at_nhds_within_of_is_separable hK.measurable_set hK.is_separable hx, end @@ -243,7 +298,6 @@ section monotone variables [borel_space X] [conditionally_complete_linear_order X] [conditionally_complete_linear_order E] [order_topology X] [order_topology E] [second_countable_topology E] - {s : set X} lemma monotone_on.integrable_on_of_measure_ne_top (hmono : monotone_on f s) {a b : X} (ha : is_least s a) (hb : is_greatest s b) (hs : μ s ≠ ∞) @@ -303,3 +357,127 @@ lemma antitone.locally_integrable [is_locally_finite_measure μ] (hanti : antito hanti.dual_right.locally_integrable end monotone + +namespace measure_theory + +variables [opens_measurable_space X] {A K : set X} + +section mul + +variables [normed_ring R] [second_countable_topology_either X R] {g g' : X → R} + +lemma integrable_on.mul_continuous_on_of_subset + (hg : integrable_on g A μ) (hg' : continuous_on g' K) + (hA : measurable_set A) (hK : is_compact K) (hAK : A ⊆ K) : + integrable_on (λ x, g x * g' x) A μ := +begin + rcases is_compact.exists_bound_of_continuous_on hK hg' with ⟨C, hC⟩, + rw [integrable_on, ← mem_ℒp_one_iff_integrable] at hg ⊢, + have : ∀ᵐ x ∂(μ.restrict A), ‖g x * g' x‖ ≤ C * ‖g x‖, + { filter_upwards [ae_restrict_mem hA] with x hx, + refine (norm_mul_le _ _).trans _, + rw mul_comm, + apply mul_le_mul_of_nonneg_right (hC x (hAK hx)) (norm_nonneg _), }, + exact mem_ℒp.of_le_mul hg (hg.ae_strongly_measurable.mul $ + (hg'.mono hAK).ae_strongly_measurable hA) this, +end + +lemma integrable_on.mul_continuous_on [t2_space X] + (hg : integrable_on g K μ) (hg' : continuous_on g' K) (hK : is_compact K) : + integrable_on (λ x, g x * g' x) K μ := +hg.mul_continuous_on_of_subset hg' hK.measurable_set hK (subset.refl _) + +lemma integrable_on.continuous_on_mul_of_subset + (hg : continuous_on g K) (hg' : integrable_on g' A μ) + (hK : is_compact K) (hA : measurable_set A) (hAK : A ⊆ K) : + integrable_on (λ x, g x * g' x) A μ := +begin + rcases is_compact.exists_bound_of_continuous_on hK hg with ⟨C, hC⟩, + rw [integrable_on, ← mem_ℒp_one_iff_integrable] at hg' ⊢, + have : ∀ᵐ x ∂(μ.restrict A), ‖g x * g' x‖ ≤ C * ‖g' x‖, + { filter_upwards [ae_restrict_mem hA] with x hx, + refine (norm_mul_le _ _).trans _, + apply mul_le_mul_of_nonneg_right (hC x (hAK hx)) (norm_nonneg _), }, + exact mem_ℒp.of_le_mul hg' (((hg.mono hAK).ae_strongly_measurable hA).mul + hg'.ae_strongly_measurable) this, +end + +lemma integrable_on.continuous_on_mul [t2_space X] + (hg : continuous_on g K) (hg' : integrable_on g' K μ) (hK : is_compact K) : + integrable_on (λ x, g x * g' x) K μ := +hg'.continuous_on_mul_of_subset hg hK hK.measurable_set subset.rfl + +end mul + +section smul + +variables {𝕜 : Type*} [normed_field 𝕜] [normed_space 𝕜 E] + +lemma integrable_on.continuous_on_smul [t2_space X] [second_countable_topology_either X 𝕜] + {g : X → E} (hg : integrable_on g K μ) {f : X → 𝕜} (hf : continuous_on f K) (hK : is_compact K) : + integrable_on (λ x, f x • g x) K μ := +begin + rw [integrable_on, ←integrable_norm_iff], + { simp_rw norm_smul, + refine integrable_on.continuous_on_mul _ hg.norm hK, + exact continuous_norm.comp_continuous_on hf }, + { exact (hf.ae_strongly_measurable hK.measurable_set).smul hg.1 } +end + +lemma integrable_on.smul_continuous_on [t2_space X] [second_countable_topology_either X E] + {f : X → 𝕜} (hf : integrable_on f K μ) {g : X → E} (hg : continuous_on g K) (hK : is_compact K) : + integrable_on (λ x, f x • g x) K μ := +begin + rw [integrable_on, ←integrable_norm_iff], + { simp_rw norm_smul, + refine integrable_on.mul_continuous_on hf.norm _ hK, + exact continuous_norm.comp_continuous_on hg }, + { exact hf.1.smul (hg.ae_strongly_measurable hK.measurable_set) } +end + +end smul + +namespace locally_integrable_on + +lemma continuous_on_mul [locally_compact_space X] [t2_space X] [normed_ring R] + [second_countable_topology_either X R] + {f g : X → R} {s : set X} + (hf : locally_integrable_on f s μ) (hg : continuous_on g s) (hs : is_open s) : + locally_integrable_on (λ x, g x * f x) s μ := +begin + rw measure_theory.locally_integrable_on_iff (or.inr hs) at hf ⊢, + exact λ k hk_sub hk_c, (hf k hk_sub hk_c).continuous_on_mul (hg.mono hk_sub) hk_c +end + +lemma mul_continuous_on [locally_compact_space X] [t2_space X] [normed_ring R] + [second_countable_topology_either X R] {f g : X → R} {s : set X} + (hf : locally_integrable_on f s μ) (hg : continuous_on g s) (hs : is_open s) : + locally_integrable_on (λ x, f x * g x) s μ := +begin + rw measure_theory.locally_integrable_on_iff (or.inr hs) at hf ⊢, + exact λ k hk_sub hk_c, (hf k hk_sub hk_c).mul_continuous_on (hg.mono hk_sub) hk_c +end + +lemma continuous_on_smul [locally_compact_space X] [t2_space X] + {𝕜 : Type*} [normed_field 𝕜] [second_countable_topology_either X 𝕜] [normed_space 𝕜 E] + {f : X → E} {g : X → 𝕜} {s : set X} (hs : is_open s) + (hf : locally_integrable_on f s μ) (hg : continuous_on g s) : + locally_integrable_on (λ x, g x • f x) s μ := +begin + rw measure_theory.locally_integrable_on_iff (or.inr hs) at hf ⊢, + exact λ k hk_sub hk_c, (hf k hk_sub hk_c).continuous_on_smul (hg.mono hk_sub) hk_c +end + +lemma smul_continuous_on [locally_compact_space X] [t2_space X] + {𝕜 : Type*} [normed_field 𝕜] [second_countable_topology_either X E] [normed_space 𝕜 E] + {f : X → 𝕜} {g : X → E} {s : set X} (hs : is_open s) + (hf : locally_integrable_on f s μ) (hg : continuous_on g s) : + locally_integrable_on (λ x, f x • g x) s μ := +begin + rw measure_theory.locally_integrable_on_iff (or.inr hs) at hf ⊢, + exact λ k hk_sub hk_c, (hf k hk_sub hk_c).smul_continuous_on (hg.mono hk_sub) hk_c +end + +end locally_integrable_on + +end measure_theory diff --git a/src/measure_theory/integral/integrable_on.lean b/src/measure_theory/integral/integrable_on.lean index a1edc3eddc9b1..fb5f30403ec4c 100644 --- a/src/measure_theory/integral/integrable_on.lean +++ b/src/measure_theory/integral/integrable_on.lean @@ -114,21 +114,26 @@ lemma integrable_on.congr_set_ae (h : integrable_on f t μ) (hst : s =ᵐ[μ] t) integrable_on f s μ := h.mono_set_ae hst.le -lemma integrable_on.congr_fun' (h : integrable_on f s μ) (hst : f =ᵐ[μ.restrict s] g) : +lemma integrable_on.congr_fun_ae (h : integrable_on f s μ) (hst : f =ᵐ[μ.restrict s] g) : integrable_on g s μ := integrable.congr h hst +lemma integrable_on_congr_fun_ae (hst : f =ᵐ[μ.restrict s] g) : + integrable_on f s μ ↔ integrable_on g s μ := +⟨λ h, h.congr_fun_ae hst, λ h, h.congr_fun_ae hst.symm⟩ + lemma integrable_on.congr_fun (h : integrable_on f s μ) (hst : eq_on f g s) (hs : measurable_set s) : integrable_on g s μ := -h.congr_fun' ((ae_restrict_iff' hs).2 (eventually_of_forall hst)) +h.congr_fun_ae ((ae_restrict_iff' hs).2 (eventually_of_forall hst)) + +lemma integrable_on_congr_fun (hst : eq_on f g s) (hs : measurable_set s) : + integrable_on f s μ ↔ integrable_on g s μ := +⟨λ h, h.congr_fun hst hs, λ h, h.congr_fun hst.symm hs⟩ lemma integrable.integrable_on (h : integrable f μ) : integrable_on f s μ := h.mono_measure $ measure.restrict_le_self -lemma integrable.integrable_on' (h : integrable f (μ.restrict s)) : integrable_on f s μ := -h - lemma integrable_on.restrict (h : integrable_on f s μ) (hs : measurable_set s) : integrable_on f s (μ.restrict t) := by { rw [integrable_on, measure.restrict_restrict hs], exact h.mono_set (inter_subset_left _ _) } diff --git a/src/measure_theory/integral/interval_integral.lean b/src/measure_theory/integral/interval_integral.lean index 81dd304d6c1b2..6b96f2e9d66d4 100644 --- a/src/measure_theory/integral/interval_integral.lean +++ b/src/measure_theory/integral/interval_integral.lean @@ -616,8 +616,9 @@ lemma integral_undef (h : ¬ interval_integrable f μ a b) : ∫ x in a..b, f x ∂μ = 0 := by cases le_total a b with hab hab; simp only [integral_of_le, integral_of_ge, hab, neg_eq_zero]; - refine integral_undef (not_imp_not.mpr integrable.integrable_on' _); - simpa [hab] using not_and_distrib.mp h + refine integral_undef (not_imp_not.mpr _ h); + simpa only [hab, Ioc_eq_empty_of_le, integrable_on_empty, not_true, false_or, or_false] + using not_and_distrib.mp h lemma interval_integrable_of_integral_ne_zero {a b : ℝ} {f : ℝ → E} {μ : measure ℝ} (h : ∫ x in a..b, f x ∂μ ≠ 0) : diff --git a/src/measure_theory/integral/layercake.lean b/src/measure_theory/integral/layercake.lean index f95873332e603..6cc9602fe672a 100644 --- a/src/measure_theory/integral/layercake.lean +++ b/src/measure_theory/integral/layercake.lean @@ -170,7 +170,7 @@ begin have g_eq_G_on : ∀ t, g =ᵐ[volume.restrict (Ioc 0 t)] G, from λ t, ae_mono (measure.restrict_mono Ioc_subset_Ioi_self le_rfl) g_eq_G, have G_intble : ∀ t > 0, interval_integrable G volume 0 t, - { refine λ t t_pos, ⟨integrable_on.congr_fun' (g_intble t t_pos).1 (g_eq_G_on t), _⟩, + { refine λ t t_pos, ⟨(g_intble t t_pos).1.congr_fun_ae (g_eq_G_on t), _⟩, rw Ioc_eq_empty_of_le t_pos.lt.le, exact integrable_on_empty, }, have eq₁ : ∫⁻ t in Ioi 0, μ {a : α | t ≤ f a} * ennreal.of_real (g t) diff --git a/src/measure_theory/integral/set_integral.lean b/src/measure_theory/integral/set_integral.lean index 85908862cbf86..a5d4d950ca8aa 100644 --- a/src/measure_theory/integral/set_integral.lean +++ b/src/measure_theory/integral/set_integral.lean @@ -292,7 +292,7 @@ lemma integral_union_eq_left_of_ae (ht_eq : ∀ᵐ x ∂(μ.restrict t), f x = 0 ∫ x in (s ∪ t), f x ∂μ = ∫ x in s, f x ∂μ := begin have ht : integrable_on f t μ, - { apply integrable_on.congr_fun' integrable_on_zero, symmetry, exact ht_eq }, + { apply integrable_on_zero.congr_fun_ae, symmetry, exact ht_eq }, by_cases H : integrable_on f (s ∪ t) μ, swap, { rw [integral_undef H, integral_undef], simpa [integrable_on_union, ht] using H }, let f' := H.1.mk f, @@ -301,7 +301,7 @@ begin ... = ∫ x in s, f' x ∂μ : begin apply integral_union_eq_left_of_ae_aux _ H.1.strongly_measurable_mk - (H.congr_fun' H.1.ae_eq_mk), + (H.congr_fun_ae H.1.ae_eq_mk), filter_upwards [ht_eq, ae_mono (measure.restrict_mono (subset_union_right s t) le_rfl) H.1.ae_eq_mk] with x hx h'x, rw [← h'x, hx] From 7aac545b979fe3ffab5c93b833129e5c8d826296 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 4 Apr 2023 07:41:57 +0000 Subject: [PATCH 0768/1291] chore(linear_algebra/dimension): remove unecessary `nontrivial` argument (#18725) This is implied by another hypothesis via `nontrivial_of_invariant_basis_number` --- src/algebra/linear_recurrence.lean | 2 +- src/linear_algebra/dimension.lean | 13 +++++++------ 2 files changed, 8 insertions(+), 7 deletions(-) diff --git a/src/algebra/linear_recurrence.lean b/src/algebra/linear_recurrence.lean index c26e070cabc6f..14461e059c4cd 100644 --- a/src/algebra/linear_recurrence.lean +++ b/src/algebra/linear_recurrence.lean @@ -177,7 +177,7 @@ variables {α : Type*} [comm_ring α] [strong_rank_condition α] (E : linear_rec lemma sol_space_dim : module.rank α E.sol_space = E.order := begin letI := nontrivial_of_invariant_basis_number α, - exact @dim_fin_fun α _ _ _ E.order ▸ E.to_init.dim_eq + exact @dim_fin_fun α _ _ E.order ▸ E.to_init.dim_eq end end strong_rank_condition diff --git a/src/linear_algebra/dimension.lean b/src/linear_algebra/dimension.lean index b4991d5081098..f503145501981 100644 --- a/src/linear_algebra/dimension.lean +++ b/src/linear_algebra/dimension.lean @@ -1005,9 +1005,10 @@ variables [∀i, add_comm_group (φ i)] [∀i, module K (φ i)] [∀i, module.fr open linear_map -lemma dim_pi [nontrivial K] [finite η] : +lemma dim_pi [finite η] : module.rank K (Πi, φ i) = cardinal.sum (λi, module.rank K (φ i)) := begin + haveI := nontrivial_of_invariant_basis_number K, casesI nonempty_fintype η, let b := λ i, (module.free.exists_basis K (φ i)).some.2, let this : basis (Σ j, _) K (Π j, φ j) := pi.basis b, @@ -1018,20 +1019,20 @@ end variable [fintype η] -lemma dim_fun {V η : Type u} [nontrivial K] [fintype η] [add_comm_group V] [module K V] +lemma dim_fun {V η : Type u} [fintype η] [add_comm_group V] [module K V] [module.free K V] : module.rank K (η → V) = fintype.card η * module.rank K V := by rw [dim_pi, cardinal.sum_const', cardinal.mk_fintype] -lemma dim_fun_eq_lift_mul [nontrivial K] : +lemma dim_fun_eq_lift_mul : module.rank K (η → V) = (fintype.card η : cardinal.{max u₁' v}) * cardinal.lift.{u₁'} (module.rank K V) := by rw [dim_pi, cardinal.sum_const, cardinal.mk_fintype, cardinal.lift_nat_cast] -lemma dim_fun' [nontrivial K] : module.rank K (η → K) = fintype.card η := +lemma dim_fun' : module.rank K (η → K) = fintype.card η := by rw [dim_fun_eq_lift_mul, dim_self, cardinal.lift_one, mul_one, cardinal.nat_cast_inj] -lemma dim_fin_fun [nontrivial K] (n : ℕ) : module.rank K (fin n → K) = n := +lemma dim_fin_fun (n : ℕ) : module.rank K (fin n → K) = n := by simp [dim_fun'] end fintype @@ -1053,7 +1054,7 @@ begin by simp, have hn := cardinal.lift_inj.{v u}.2 hn, rw this at hn, - rw ←@dim_fin_fun K _ _ _ n at hn, + rw ←@dim_fin_fun K _ _ n at hn, haveI : module.free K (fin n → K) := module.free.pi _ _, exact classical.choice (nonempty_linear_equiv_of_lift_dim_eq hn), end From 0ce07501ebb0dee26dde9f0805380cf68975dd24 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 4 Apr 2023 09:42:43 +0000 Subject: [PATCH 0769/1291] chore(linear_algebra/finite_dimensional): generalize from Type to Sort (#18723) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The replacement of `fintype` with `finite` means there's no longer a need to handle these cases (added in #12877) separately. I verified that `{ι : Sort*}` is not being inferred as `{ι : Sort (u + 1)}`. --- src/linear_algebra/finite_dimensional.lean | 20 ++++---------------- 1 file changed, 4 insertions(+), 16 deletions(-) diff --git a/src/linear_algebra/finite_dimensional.lean b/src/linear_algebra/finite_dimensional.lean index 74d8af2c6886e..44e3ce1967570 100644 --- a/src/linear_algebra/finite_dimensional.lean +++ b/src/linear_algebra/finite_dimensional.lean @@ -688,27 +688,15 @@ begin end /-- The submodule generated by a supremum of finite dimensional submodules, indexed by a finite -type is finite-dimensional. -/ -instance finite_dimensional_supr {ι : Type*} [_root_.finite ι] (S : ι → submodule K V) +sort is finite-dimensional. -/ +instance finite_dimensional_supr {ι : Sort*} [_root_.finite ι] (S : ι → submodule K V) [Π i, finite_dimensional K (S i)] : finite_dimensional K ↥(⨆ i, S i) := begin - casesI nonempty_fintype ι, - rw ←finset.sup_univ_eq_supr, + casesI nonempty_fintype (plift ι), + rw [←supr_plift_down, ← finset.sup_univ_eq_supr], exact submodule.finite_dimensional_finset_sup _ _, end -/-- The submodule generated by a supremum indexed by a proposition is finite-dimensional if -the submodule is. -/ -instance finite_dimensional_supr_prop {P : Prop} (S : P → submodule K V) - [Π h, finite_dimensional K (S h)] : finite_dimensional K ↥(⨆ h, S h) := -begin - by_cases hp : P, - { rw supr_pos hp, - apply_instance }, - { rw supr_neg hp, - 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 From 4cf7ca0e69e048b006674cf4499e5c7d296a89e0 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 4 Apr 2023 11:37:56 +0000 Subject: [PATCH 0770/1291] chore(linear_algebra/free_module/finite/rank): move lemmas from `module.free` to `finite_dimensional` (#18733) The lemmas about finite-dimensional spaces are currently scattered between namespaces. This commit mostly addresses the confusion by renaming all the `module.free.finrank_*` lemmas to `finite_dimensional.finrank_*`. This rename makes it apparent that `finrank_eq_dim` and `finrank_eq_rank` are duplicates; though it seems that for performances reasons it's still useful in one or two places to keep both. [Zulip thread](https://leanprover.zulipchat.com/#narrow/stream/113488-general/topic/naming.3A.20finrank.20and.20rank.20lemmas/near/346701602) --- src/data/matrix/rank.lean | 6 +++--- src/linear_algebra/finite_dimensional.lean | 11 ++++------- src/linear_algebra/free_module/finite/rank.lean | 8 ++++---- src/linear_algebra/trace.lean | 2 +- src/number_theory/ramification_inertia.lean | 2 +- src/ring_theory/dedekind_domain/integral_closure.lean | 3 ++- 6 files changed, 15 insertions(+), 17 deletions(-) diff --git a/src/data/matrix/rank.lean b/src/data/matrix/rank.lean index 0e4d9dcf15364..2bc99be10f82d 100644 --- a/src/data/matrix/rank.lean +++ b/src/data/matrix/rank.lean @@ -38,7 +38,7 @@ variables (A : matrix m n K) noncomputable def rank : ℕ := finrank K A.to_lin'.range @[simp] lemma rank_one : rank (1 : matrix n n K) = fintype.card n := -by rw [rank, to_lin'_one, linear_map.range_id, finrank_top, module.free.finrank_pi] +by rw [rank, to_lin'_one, linear_map.range_id, finrank_top, finrank_pi] @[simp] lemma rank_zero : rank (0 : matrix n n K) = 0 := by rw [rank, linear_equiv.map_zero, linear_map.range_zero, finrank_bot] @@ -46,7 +46,7 @@ by rw [rank, linear_equiv.map_zero, linear_map.range_zero, finrank_bot] lemma rank_le_card_width : A.rank ≤ fintype.card n := begin convert le_of_add_le_left (A.to_lin'.finrank_range_add_finrank_ker).le, - exact (module.free.finrank_pi K).symm, + exact (finrank_pi K).symm, end lemma rank_le_width {m n : ℕ} (A : matrix (fin m) (fin n) K) : A.rank ≤ n := @@ -96,7 +96,7 @@ begin end lemma rank_le_card_height : A.rank ≤ fintype.card m := -(submodule.finrank_le _).trans (module.free.finrank_pi K).le +(submodule.finrank_le _).trans (finrank_pi K).le omit m_fin diff --git a/src/linear_algebra/finite_dimensional.lean b/src/linear_algebra/finite_dimensional.lean index 44e3ce1967570..653626d9735df 100644 --- a/src/linear_algebra/finite_dimensional.lean +++ b/src/linear_algebra/finite_dimensional.lean @@ -168,13 +168,10 @@ module.finite.of_surjective (submodule.mkq S) $ surjective_quot_mk _ variables (K V) /-- In a finite-dimensional space, its dimension (seen as a cardinal) coincides with its -`finrank`. -/ +`finrank`. This is a copy of `finrank_eq_rank _ _` which creates easier typeclass searches. -/ lemma finrank_eq_dim [finite_dimensional K V] : (finrank K V : cardinal.{v}) = module.rank K V := -begin - letI : is_noetherian K V := iff_fg.2 infer_instance, - rw [finrank, cast_to_nat_of_lt_aleph_0 (dim_lt_aleph_0 K V)] -end +finrank_eq_rank _ _ variables {K V} lemma finrank_of_infinite_dimensional (h : ¬finite_dimensional K V) : finrank K V = 0 := @@ -252,7 +249,7 @@ lemma cardinal_mk_le_finrank_of_linear_independent #ι ≤ finrank K V := begin rw ← lift_le.{_ (max v w)}, - simpa [← finrank_eq_dim, -module.free.finrank_eq_rank] using + simpa [← finrank_eq_dim, -finrank_eq_rank] using cardinal_lift_le_dim_of_linear_independent.{_ _ _ (max v w)} h end @@ -371,7 +368,7 @@ end /-- Pushforwards of finite-dimensional submodules have a smaller finrank. -/ lemma finrank_map_le (f : V →ₗ[K] V₂) (p : submodule K V) [finite_dimensional K p] : finrank K (p.map f) ≤ finrank K p := -by simpa [← finrank_eq_dim, -module.free.finrank_eq_rank] using lift_dim_map_le f p +by simpa [← finrank_eq_dim, -finrank_eq_rank] using lift_dim_map_le f p variable {K} diff --git a/src/linear_algebra/free_module/finite/rank.lean b/src/linear_algebra/free_module/finite/rank.lean index d7d0349d93e9c..8edb37749b494 100644 --- a/src/linear_algebra/free_module/finite/rank.lean +++ b/src/linear_algebra/free_module/finite/rank.lean @@ -16,8 +16,7 @@ This is a basic API for the rank of finite free modules. -/ ---TODO: `linear_algebra/finite_dimensional` should import this file, and a lot of results should ---be moved here. +--TODO: many results from `linear_algebra/finite_dimensional` should be moved here. universes u v w @@ -27,7 +26,8 @@ open_locale tensor_product direct_sum big_operators cardinal open cardinal finite_dimensional fintype -namespace module.free +namespace finite_dimensional +open module.free section ring @@ -110,4 +110,4 @@ by { simp [finrank] } end comm_ring -end module.free +end finite_dimensional diff --git a/src/linear_algebra/trace.lean b/src/linear_algebra/trace.lean index 1eff304e71d48..c164ee18444f2 100644 --- a/src/linear_algebra/trace.lean +++ b/src/linear_algebra/trace.lean @@ -165,7 +165,7 @@ trace_eq_contract_of_basis' (module.free.choose_basis R M) @[simp] theorem trace_one : trace R M 1 = (finrank R M : R) := begin have b := module.free.choose_basis R M, - rw [trace_eq_matrix_trace R b, to_matrix_one, module.free.finrank_eq_card_choose_basis_index], + rw [trace_eq_matrix_trace R b, to_matrix_one, finrank_eq_card_choose_basis_index], simp, end diff --git a/src/number_theory/ramification_inertia.lean b/src/number_theory/ramification_inertia.lean index 38a4402e6e9bb..2e66725f662bc 100644 --- a/src/number_theory/ramification_inertia.lean +++ b/src/number_theory/ramification_inertia.lean @@ -815,7 +815,7 @@ begin finrank (R ⧸ p) (S ⧸ (P : ideal S)^(e P)) : _ ... = finrank (R ⧸ p) (Π P : (factors (map (algebra_map R S) p)).to_finset, (S ⧸ (P : ideal S)^(e P))) : - (module.free.finrank_pi_fintype (R ⧸ p)).symm + (finrank_pi_fintype (R ⧸ p)).symm ... = finrank (R ⧸ p) (S ⧸ map (algebra_map R S) p) : _ ... = finrank K L : _, { rw ← finset.sum_attach, diff --git a/src/ring_theory/dedekind_domain/integral_closure.lean b/src/ring_theory/dedekind_domain/integral_closure.lean index 8d33ce8a496d2..8116312cb9f07 100644 --- a/src/ring_theory/dedekind_domain/integral_closure.lean +++ b/src/ring_theory/dedekind_domain/integral_closure.lean @@ -233,7 +233,8 @@ begin haveI : is_localization (algebra.algebra_map_submonoid C A⁰) L := is_integral_closure.is_localization A K L C, let b := basis.localization_localization K A⁰ L (module.free.choose_basis A C), - rw [module.free.finrank_eq_card_choose_basis_index, finite_dimensional.finrank_eq_card_basis b], + rw [finite_dimensional.finrank_eq_card_choose_basis_index, + finite_dimensional.finrank_eq_card_basis b], end variables {A K} From be2ac64be57e8319fcd5c5547f3a8d9412daf5ec Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 4 Apr 2023 11:37:57 +0000 Subject: [PATCH 0771/1291] chore(linear_algebra/dimension): rename `rank` to `linear_map.rank` (#18734) It's a bit weird having `rank` in the root namespace, defined in terms of `module.rank` which isn't. This moves all the lemmas into a single block at the end of the file, as they were previously spread through the file. --- docs/undergrad.yaml | 2 +- .../normed_space/finite_dimension.lean | 4 +- src/linear_algebra/dimension.lean | 112 ++++++++++-------- 3 files changed, 63 insertions(+), 55 deletions(-) diff --git a/docs/undergrad.yaml b/docs/undergrad.yaml index 134f166f3b276..b9fe40e3e405a 100644 --- a/docs/undergrad.yaml +++ b/docs/undergrad.yaml @@ -28,7 +28,7 @@ Linear algebra: Finite-dimensional vector spaces: finite-dimensionality : 'finite_dimensional' isomorphism with $K^n$: 'basis.equiv_fun' - rank of a linear map: 'rank' + rank of a linear map: 'linear_map.rank' rank of a set of vectors: 'set.finrank' rank of a system of linear equations: 'https://www.math.tamu.edu/~fnarc/psfiles/rank2005.pdf' isomorphism with bidual: 'module.eval_equiv' diff --git a/src/analysis/normed_space/finite_dimension.lean b/src/analysis/normed_space/finite_dimension.lean index 8ef4d84dab195..2aeddb53260b9 100644 --- a/src/analysis/normed_space/finite_dimension.lean +++ b/src/analysis/normed_space/finite_dimension.lean @@ -244,9 +244,9 @@ lemma is_open_set_of_linear_independent {ι : Type*} [finite ι] : is_open {f : ι → E | linear_independent 𝕜 f} := is_open_iff_mem_nhds.2 $ λ f, linear_independent.eventually -lemma is_open_set_of_nat_le_rank (n : ℕ) : is_open {f : E →L[𝕜] F | ↑n ≤ rank (f : E →ₗ[𝕜] F)} := +lemma is_open_set_of_nat_le_rank (n : ℕ) : is_open {f : E →L[𝕜] F | ↑n ≤ (f : E →ₗ[𝕜] F).rank } := begin - simp only [le_rank_iff_exists_linear_independent_finset, set_of_exists, ← exists_prop], + simp only [linear_map.le_rank_iff_exists_linear_independent_finset, set_of_exists, ← exists_prop], refine is_open_bUnion (λ t ht, _), have : continuous (λ f : E →L[𝕜] F, (λ x : (t : set E), f x)), from continuous_pi (λ x, (continuous_linear_map.apply 𝕜 F (x : E)).continuous), diff --git a/src/linear_algebra/dimension.lean b/src/linear_algebra/dimension.lean index f503145501981..894f946ba1206 100644 --- a/src/linear_algebra/dimension.lean +++ b/src/linear_algebra/dimension.lean @@ -1169,61 +1169,10 @@ end end division_ring -section rank - -section -variables [ring K] [add_comm_group V] [module K V] [add_comm_group V₁] [module K V₁] -variables [add_comm_group V'] [module K V'] - -/-- `rank f` is the rank of a `linear_map f`, defined as the dimension of `f.range`. -/ -def rank (f : V →ₗ[K] V') : cardinal := module.rank K f.range - -lemma rank_le_range (f : V →ₗ[K] V₁) : rank f ≤ module.rank K V₁ := -dim_submodule_le _ - -@[simp] lemma rank_zero [nontrivial K] : rank (0 : V →ₗ[K] V') = 0 := -by rw [rank, linear_map.range_zero, dim_bot] - -variables [add_comm_group V''] [module K V''] - -lemma rank_comp_le1 (g : V →ₗ[K] V') (f : V' →ₗ[K] V'') : rank (f.comp g) ≤ rank f := -begin - refine dim_le_of_submodule _ _ _, - rw [linear_map.range_comp], - exact linear_map.map_le_range, -end - -variables [add_comm_group V'₁] [module K V'₁] - -lemma rank_comp_le2 (g : V →ₗ[K] V') (f : V' →ₗ[K] V'₁) : rank (f.comp g) ≤ rank g := -by rw [rank, rank, linear_map.range_comp]; exact dim_map_le _ _ - -end - -end rank - section division_ring variables [division_ring K] [add_comm_group V] [module K V] [add_comm_group V₁] [module K V₁] variables [add_comm_group V'] [module K V'] -lemma rank_le_domain (f : V →ₗ[K] V₁) : rank f ≤ module.rank K V := -by { rw [← dim_range_add_dim_ker f], exact self_le_add_right _ _ } - -lemma rank_add_le (f g : V →ₗ[K] V') : rank (f + g) ≤ rank f + rank g := -calc rank (f + g) ≤ module.rank K (f.range ⊔ g.range : submodule K V') : - begin - refine dim_le_of_submodule _ _ _, - exact (linear_map.range_le_iff_comap.2 $ eq_top_iff'.2 $ - assume x, show f x + g x ∈ (f.range ⊔ g.range : submodule K V'), from - mem_sup.2 ⟨_, ⟨x, rfl⟩, _, ⟨x, rfl⟩, rfl⟩) - end - ... ≤ rank f + rank g : dim_add_le_dim_add_dim _ _ - -lemma rank_finset_sum_le {η} (s : finset η) (f : η → V →ₗ[K] V') : - rank (∑ d in s, f d) ≤ ∑ d in s, rank (f d) := -@finset.sum_hom_rel _ _ _ _ _ (λa b, rank a ≤ b) f (λ d, rank (f d)) s (le_of_eq rank_zero) - (λ i g c h, le_trans (rank_add_le _ _) (add_le_add_left h _)) - /-- The `ι` indexed basis on `V`, where `ι` is an empty type and `V` is zero-dimensional. See also `finite_dimensional.fin_basis`. @@ -1355,6 +1304,65 @@ lemma module.rank_le_one_iff_top_is_principal : module.rank K V ≤ 1 ↔ (⊤ : submodule K V).is_principal := by rw [← submodule.rank_le_one_iff_is_principal, dim_top] +end division_ring + +end module + +/-! ### The rank of a linear map -/ + +namespace linear_map + +section ring +variables [ring K] [add_comm_group V] [module K V] [add_comm_group V₁] [module K V₁] +variables [add_comm_group V'] [module K V'] + +/-- `rank f` is the rank of a `linear_map` `f`, defined as the dimension of `f.range`. -/ +def rank (f : V →ₗ[K] V') : cardinal := module.rank K f.range + +lemma rank_le_range (f : V →ₗ[K] V₁) : rank f ≤ module.rank K V₁ := +dim_submodule_le _ + +@[simp] lemma rank_zero [nontrivial K] : rank (0 : V →ₗ[K] V') = 0 := +by rw [rank, linear_map.range_zero, dim_bot] + +variables [add_comm_group V''] [module K V''] + +lemma rank_comp_le1 (g : V →ₗ[K] V') (f : V' →ₗ[K] V'') : rank (f.comp g) ≤ rank f := +begin + refine dim_le_of_submodule _ _ _, + rw [linear_map.range_comp], + exact linear_map.map_le_range, +end + +variables [add_comm_group V'₁] [module K V'₁] + +lemma rank_comp_le2 (g : V →ₗ[K] V') (f : V' →ₗ[K] V'₁) : rank (f.comp g) ≤ rank g := +by rw [rank, rank, linear_map.range_comp]; exact dim_map_le _ _ + +end ring + +section division_ring +variables [division_ring K] [add_comm_group V] [module K V] [add_comm_group V₁] [module K V₁] +variables [add_comm_group V'] [module K V'] + +lemma rank_le_domain (f : V →ₗ[K] V₁) : rank f ≤ module.rank K V := +by { rw [← dim_range_add_dim_ker f], exact self_le_add_right _ _ } + +lemma rank_add_le (f g : V →ₗ[K] V') : rank (f + g) ≤ rank f + rank g := +calc rank (f + g) ≤ module.rank K (f.range ⊔ g.range : submodule K V') : + begin + refine dim_le_of_submodule _ _ _, + exact (linear_map.range_le_iff_comap.2 $ eq_top_iff'.2 $ + assume x, show f x + g x ∈ (f.range ⊔ g.range : submodule K V'), from + mem_sup.2 ⟨_, ⟨x, rfl⟩, _, ⟨x, rfl⟩, rfl⟩) + end + ... ≤ rank f + rank g : dim_add_le_dim_add_dim _ _ + +lemma rank_finset_sum_le {η} (s : finset η) (f : η → V →ₗ[K] V') : + rank (∑ d in s, f d) ≤ ∑ d in s, rank (f d) := +@finset.sum_hom_rel _ _ _ _ _ (λa b, rank a ≤ b) f (λ d, rank (f d)) s (le_of_eq rank_zero) + (λ i g c h, le_trans (rank_add_le _ _) (add_le_add_left h _)) + lemma le_rank_iff_exists_linear_independent {c : cardinal} {f : V →ₗ[K] V'} : c ≤ rank f ↔ ∃ s : set V, cardinal.lift.{v'} (#s) = cardinal.lift.{v} c ∧ @@ -1391,4 +1399,4 @@ end end division_ring -end module +end linear_map From d524d0a578db1146460c1aca35bb5db68466347a Mon Sep 17 00:00:00 2001 From: Moritz Doll Date: Tue, 4 Apr 2023 13:35:11 +0000 Subject: [PATCH 0772/1291] feat(analysis/calculus/iterated_deriv): equality of norms of iterated derivative (#18728) --- src/analysis/calculus/iterated_deriv.lean | 8 ++++++++ 1 file changed, 8 insertions(+) diff --git a/src/analysis/calculus/iterated_deriv.lean b/src/analysis/calculus/iterated_deriv.lean index 0282cb9a4197a..2d1050774167c 100644 --- a/src/analysis/calculus/iterated_deriv.lean +++ b/src/analysis/calculus/iterated_deriv.lean @@ -96,6 +96,10 @@ begin simp end +lemma norm_iterated_fderiv_within_eq_norm_iterated_deriv_within : + ‖iterated_fderiv_within 𝕜 n f s x‖ = ‖iterated_deriv_within n f s x‖ := +by rw [iterated_deriv_within_eq_equiv_comp, linear_isometry_equiv.norm_map] + @[simp] lemma iterated_deriv_within_zero : iterated_deriv_within 0 f s = f := by { ext x, simp [iterated_deriv_within] } @@ -222,6 +226,10 @@ lemma iterated_fderiv_apply_eq_iterated_deriv_mul_prod {m : (fin n) → 𝕜} : (iterated_fderiv 𝕜 n f x : ((fin n) → 𝕜) → F) m = (∏ i, m i) • iterated_deriv n f x := by { rw [iterated_deriv_eq_iterated_fderiv, ← continuous_multilinear_map.map_smul_univ], simp } +lemma norm_iterated_fderiv_eq_norm_iterated_deriv : + ‖iterated_fderiv 𝕜 n f x‖ = ‖iterated_deriv n f x‖ := +by rw [iterated_deriv_eq_equiv_comp, linear_isometry_equiv.norm_map] + @[simp] lemma iterated_deriv_zero : iterated_deriv 0 f = f := by { ext x, simp [iterated_deriv] } From 1a4df69ca1a9a0e5e26bfe12e2b92814216016d0 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 4 Apr 2023 16:43:38 +0000 Subject: [PATCH 0773/1291] feat(analysis/inner_product_space/basic): orthogonal submodules (#18705) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This adds `submodule.is_ortho U V`, with notation `U ⟂ V`, as a shorthand for `U ≤ Vᗮ`. To make this useful, this also adds about 30 lemmas of basic API. Some downstream proofs are golfed using the new API. --- src/analysis/inner_product_space/basic.lean | 1 + .../gram_schmidt_ortho.lean | 22 +-- .../inner_product_space/orthogonal.lean | 161 +++++++++++++++++- .../inner_product_space/projection.lean | 17 ++ .../inner_product_space/spectrum.lean | 4 +- src/geometry/euclidean/monge_point.lean | 11 +- 6 files changed, 193 insertions(+), 23 deletions(-) diff --git a/src/analysis/inner_product_space/basic.lean b/src/analysis/inner_product_space/basic.lean index 557e331d90c23..242bb6c7a5b9d 100644 --- a/src/analysis/inner_product_space/basic.lean +++ b/src/analysis/inner_product_space/basic.lean @@ -1885,6 +1885,7 @@ open_locale direct_sum The simple way to express this concept would be as a condition on `V : ι → submodule 𝕜 E`. We We instead implement it as a condition on a family of inner product spaces each equipped with an isometric embedding into `E`, thus making it a property of morphisms rather than subobjects. +The connection to the subobject spelling is shown in `orthogonal_family_iff_pairwise`. This definition is less lightweight, but allows for better definitional properties when the inner product space structure on each of the submodules is important -- for example, when considering diff --git a/src/analysis/inner_product_space/gram_schmidt_ortho.lean b/src/analysis/inner_product_space/gram_schmidt_ortho.lean index da9d09de3772d..a46b0f5b0a442 100644 --- a/src/analysis/inner_product_space/gram_schmidt_ortho.lean +++ b/src/analysis/inner_product_space/gram_schmidt_ortho.lean @@ -181,17 +181,15 @@ begin apply finset.sum_eq_zero, intros j hj, rw coe_eq_zero, - suffices : span 𝕜 (f '' set.Iic j) ≤ (𝕜 ∙ f i)ᗮ, + suffices : span 𝕜 (f '' set.Iic j) ⟂ 𝕜 ∙ f i, { apply orthogonal_projection_mem_subspace_orthogonal_complement_eq_zero, rw mem_orthogonal_singleton_iff_inner_left, rw ←mem_orthogonal_singleton_iff_inner_right, exact this (gram_schmidt_mem_span 𝕜 f (le_refl j)) }, - rw span_le, - rintros - ⟨k, hk, rfl⟩, - rw [set_like.mem_coe, mem_orthogonal_singleton_iff_inner_left], + rw is_ortho_span, + rintros u ⟨k, hk, rfl⟩ v (rfl : v = f i), apply hf, - refine (lt_of_le_of_lt hk _).ne, - simpa using hj }, + exact (lt_of_le_of_lt hk (finset.mem_Iio.mp hj)).ne }, { simp }, end @@ -358,16 +356,14 @@ lemma inner_gram_schmidt_orthonormal_basis_eq_zero {f : ι → E} {i : ι} ⟪gram_schmidt_orthonormal_basis h f i, f j⟫ = 0 := begin rw ←mem_orthogonal_singleton_iff_inner_right, - suffices : span 𝕜 (gram_schmidt_normed 𝕜 f '' Iic j) - ≤ (𝕜 ∙ gram_schmidt_orthonormal_basis h f i)ᗮ, + suffices : span 𝕜 (gram_schmidt_normed 𝕜 f '' Iic j) ⟂ 𝕜 ∙ gram_schmidt_orthonormal_basis h f i, { apply this, rw span_gram_schmidt_normed, - simpa using mem_span_gram_schmidt 𝕜 f (le_refl j) }, - rw span_le, - rintros - ⟨k, -, rfl⟩, - rw [set_like.mem_coe, mem_orthogonal_singleton_iff_inner_left], + exact mem_span_gram_schmidt 𝕜 f le_rfl }, + rw is_ortho_span, + rintros u ⟨k, hk, rfl⟩ v (rfl : v = _), by_cases hk : gram_schmidt_normed 𝕜 f k = 0, - { simp [hk] }, + { rw [hk, inner_zero_left] }, rw ← gram_schmidt_orthonormal_basis_apply h hk, have : k ≠ i, { rintros rfl, diff --git a/src/analysis/inner_product_space/orthogonal.lean b/src/analysis/inner_product_space/orthogonal.lean index e74f2e7011cab..67955868a091d 100644 --- a/src/analysis/inner_product_space/orthogonal.lean +++ b/src/analysis/inner_product_space/orthogonal.lean @@ -19,14 +19,19 @@ In this file, the `orthogonal` complement of a submodule `K` is defined, and bas Some of the more subtle results about the orthogonal complement are delayed to `analysis.inner_product_space.projection`. +See also `bilin_form.orthogonal` for orthogonality with respect to a general bilinear form. + ## Notation The orthogonal complement of a submodule `K` is denoted by `Kᗮ`. + +The proposition that two submodules are orthogonal, `submodule.is_ortho`, is denoted by `U ⟂ V`. +Note this is not the same unicode symbol as `⊥` (`has_bot`). -/ variables {𝕜 E F : Type*} [is_R_or_C 𝕜] variables [normed_add_comm_group E] [inner_product_space 𝕜 E] -variables [normed_add_comm_group F] [inner_product_space ℝ F] +variables [normed_add_comm_group F] [inner_product_space 𝕜 F] local notation `⟪`x`, `y`⟫` := @inner 𝕜 _ _ x y namespace submodule @@ -200,3 +205,157 @@ lemma orthogonal_family_self : | ff ff := absurd rfl end submodule + +@[simp] +lemma bilin_form_of_real_inner_orthogonal {E} [normed_add_comm_group E] [inner_product_space ℝ E] + (K : submodule ℝ E) : + bilin_form_of_real_inner.orthogonal K = Kᗮ := rfl + +/-! +### Orthogonality of submodules + +In this section we define `submodule.is_ortho U V`, with notation `U ⟂ V`. + +The API roughly matches that of `disjoint`. +-/ +namespace submodule + +/-- The proposition that two submodules are orthogonal. Has notation `U ⟂ V`. -/ +def is_ortho (U V : submodule 𝕜 E) : Prop := +U ≤ Vᗮ + +infix ` ⟂ `:50 := submodule.is_ortho + +lemma is_ortho_iff_le {U V : submodule 𝕜 E} : U ⟂ V ↔ U ≤ Vᗮ := iff.rfl + +@[symm] +lemma is_ortho.symm {U V : submodule 𝕜 E} (h : U ⟂ V) : V ⟂ U := +(le_orthogonal_orthogonal _).trans (orthogonal_le h) +lemma is_ortho_comm {U V : submodule 𝕜 E} : U ⟂ V ↔ V ⟂ U := ⟨is_ortho.symm, is_ortho.symm⟩ +lemma symmetric_is_ortho : symmetric (is_ortho : submodule 𝕜 E → submodule 𝕜 E → Prop) := +λ _ _, is_ortho.symm + +lemma is_ortho.inner_eq {U V : submodule 𝕜 E} (h : U ⟂ V) {u v : E} (hu : u ∈ U) (hv : v ∈ V) : + ⟪u, v⟫ = 0 := +h.symm hv _ hu + +lemma is_ortho_iff_inner_eq {U V : submodule 𝕜 E} : U ⟂ V ↔ ∀ (u ∈ U) (v ∈ V), ⟪u, v⟫ = 0 := +forall₄_congr $ λ u hu v hv, inner_eq_zero_symm + +/- TODO: generalize `submodule.map₂` to semilinear maps, so that we can state +`U ⟂ V ↔ submodule.map₂ (innerₛₗ 𝕜) U V ≤ ⊥`. -/ + +@[simp] lemma is_ortho_bot_left {V : submodule 𝕜 E} : ⊥ ⟂ V := bot_le +@[simp] lemma is_ortho_bot_right {U : submodule 𝕜 E} : U ⟂ ⊥ := is_ortho_bot_left.symm + +lemma is_ortho.mono_left {U₁ U₂ V : submodule 𝕜 E} (hU : U₂ ≤ U₁) (h : U₁ ⟂ V) : U₂ ⟂ V := +hU.trans h + +lemma is_ortho.mono_right {U V₁ V₂ : submodule 𝕜 E} (hV : V₂ ≤ V₁) (h : U ⟂ V₁) : U ⟂ V₂ := +(h.symm.mono_left hV).symm + +lemma is_ortho.mono {U₁ V₁ U₂ V₂ : submodule 𝕜 E} (hU : U₂ ≤ U₁) (hV : V₂ ≤ V₁) (h : U₁ ⟂ V₁) : + U₂ ⟂ V₂ := +(h.mono_right hV).mono_left hU + +@[simp] +lemma is_ortho_self {U : submodule 𝕜 E} : U ⟂ U ↔ U = ⊥ := +⟨λ h, eq_bot_iff.mpr $ λ x hx, inner_self_eq_zero.mp (h hx x hx), λ h, h.symm ▸ is_ortho_bot_left⟩ + +@[simp] lemma is_ortho_orthogonal_right (U : submodule 𝕜 E) : U ⟂ Uᗮ := +le_orthogonal_orthogonal _ + +@[simp] lemma is_ortho_orthogonal_left (U : submodule 𝕜 E) : Uᗮ ⟂ U := +(is_ortho_orthogonal_right U).symm + +lemma is_ortho.le {U V : submodule 𝕜 E} (h : U ⟂ V) : U ≤ Vᗮ := h +lemma is_ortho.ge {U V : submodule 𝕜 E} (h : U ⟂ V) : V ≤ Uᗮ := h.symm +@[simp] +lemma is_ortho_top_right {U : submodule 𝕜 E} : U ⟂ ⊤ ↔ U = ⊥ := +⟨λ h, eq_bot_iff.mpr $ λ x hx, inner_self_eq_zero.mp (h hx _ mem_top), + λ h, h.symm ▸ is_ortho_bot_left⟩ + +@[simp] +lemma is_ortho_top_left {V : submodule 𝕜 E} : ⊤ ⟂ V ↔ V = ⊥ := +is_ortho_comm.trans is_ortho_top_right + +/-- Orthogonal submodules are disjoint. -/ +lemma is_ortho.disjoint {U V : submodule 𝕜 E} (h : U ⟂ V) : disjoint U V := +(submodule.orthogonal_disjoint _).mono_right h.symm + +@[simp] lemma is_ortho_sup_left {U₁ U₂ V : submodule 𝕜 E} : U₁ ⊔ U₂ ⟂ V ↔ U₁ ⟂ V ∧ U₂ ⟂ V := +sup_le_iff + +@[simp] lemma is_ortho_sup_right {U V₁ V₂ : submodule 𝕜 E} : U ⟂ V₁ ⊔ V₂ ↔ U ⟂ V₁ ∧ U ⟂ V₂ := +is_ortho_comm.trans $ is_ortho_sup_left.trans $ is_ortho_comm.and is_ortho_comm + +@[simp] lemma is_ortho_Sup_left {U : set (submodule 𝕜 E)} {V : submodule 𝕜 E} : + Sup U ⟂ V ↔ ∀ Uᵢ ∈ U, Uᵢ ⟂ V := +Sup_le_iff + +@[simp] lemma is_ortho_Sup_right {U : submodule 𝕜 E} {V : set (submodule 𝕜 E)} : + U ⟂ Sup V ↔ ∀ Vᵢ ∈ V, U ⟂ Vᵢ := +is_ortho_comm.trans $ is_ortho_Sup_left.trans $ by simp_rw is_ortho_comm + +@[simp] lemma is_ortho_supr_left {ι : Sort*} {U : ι → submodule 𝕜 E} {V : submodule 𝕜 E} : + supr U ⟂ V ↔ ∀ i, U i ⟂ V := +supr_le_iff + +@[simp] lemma is_ortho_supr_right {ι : Sort*} {U : submodule 𝕜 E} {V : ι → submodule 𝕜 E} : + U ⟂ supr V ↔ ∀ i, U ⟂ V i := +is_ortho_comm.trans $ is_ortho_supr_left.trans $ by simp_rw is_ortho_comm + +@[simp] lemma is_ortho_span {s t : set E} : + span 𝕜 s ⟂ span 𝕜 t ↔ ∀ ⦃u⦄, u ∈ s → ∀ ⦃v⦄, v ∈ t → ⟪u, v⟫ = 0 := +begin + simp_rw [span_eq_supr_of_singleton_spans s, span_eq_supr_of_singleton_spans t, + is_ortho_supr_left, is_ortho_supr_right, is_ortho_iff_le, span_le, set.subset_def, + set_like.mem_coe, mem_orthogonal_singleton_iff_inner_left, set.mem_singleton_iff, forall_eq], +end + +lemma is_ortho.map (f : E →ₗᵢ[𝕜] F) {U V : submodule 𝕜 E} (h : U ⟂ V) : U.map f ⟂ V.map f := +begin + rw is_ortho_iff_inner_eq at *, + simp_rw [mem_map, forall_exists_index, and_imp, forall_apply_eq_imp_iff₂, + linear_isometry.inner_map_map], + exact h, +end + +lemma is_ortho.comap (f : E →ₗᵢ[𝕜] F) {U V : submodule 𝕜 F} (h : U ⟂ V) : U.comap f ⟂ V.comap f := +begin + rw is_ortho_iff_inner_eq at *, + simp_rw [mem_comap, ←f.inner_map_map], + intros u hu v hv, + exact h _ hu _ hv, +end + +@[simp] lemma is_ortho.map_iff (f : E ≃ₗᵢ[𝕜] F) {U V : submodule 𝕜 E} : U.map f ⟂ V.map f ↔ U ⟂ V := +⟨λ h, begin + have hf : ∀ p : submodule 𝕜 E, (p.map f).comap f.to_linear_isometry = p := + comap_map_eq_of_injective f.injective, + simpa only [hf] using h.comap f.to_linear_isometry, +end, is_ortho.map f.to_linear_isometry⟩ + +@[simp] lemma is_ortho.comap_iff (f : E ≃ₗᵢ[𝕜] F) {U V : submodule 𝕜 F} : + U.comap f ⟂ V.comap f ↔ U ⟂ V := +⟨λ h, begin + have hf : ∀ p : submodule 𝕜 F, (p.comap f).map f.to_linear_isometry = p := + map_comap_eq_of_surjective f.surjective, + simpa only [hf] using h.map f.to_linear_isometry, +end, is_ortho.comap f.to_linear_isometry⟩ + +end submodule + +lemma orthogonal_family_iff_pairwise {ι} {V : ι → submodule 𝕜 E} : + orthogonal_family 𝕜 (λ i, V i) (λ i, (V i).subtypeₗᵢ) ↔ pairwise ((⟂) on V) := +forall₃_congr $ λ i j hij, + subtype.forall.trans $ forall₂_congr $ λ x hx, subtype.forall.trans $ forall₂_congr $ λ y hy, + inner_eq_zero_symm + +alias orthogonal_family_iff_pairwise ↔ orthogonal_family.pairwise orthogonal_family.of_pairwise + +/-- Two submodules in an orthogonal family with different indices are orthogonal. -/ +lemma orthogonal_family.is_ortho {ι} {V : ι → submodule 𝕜 E} + (hV : orthogonal_family 𝕜 (λ i, V i) (λ i, (V i).subtypeₗᵢ)) {i j : ι} (hij : i ≠ j) : + V i ⟂ V j := +hV.pairwise hij diff --git a/src/analysis/inner_product_space/projection.lean b/src/analysis/inner_product_space/projection.lean index 04da7b4e7a6b1..cc4c578cbd439 100644 --- a/src/analysis/inner_product_space/projection.lean +++ b/src/analysis/inner_product_space/projection.lean @@ -785,6 +785,23 @@ lemma orthogonal_projection_mem_subspace_orthogonal_complement_eq_zero orthogonal_projection K v = 0 := by { ext, convert eq_orthogonal_projection_of_mem_orthogonal _ _; simp [hv] } +/-- The projection into `U` from an orthogonal submodule `V` is the zero map. -/ +lemma submodule.is_ortho.orthogonal_projection_comp_subtypeL {U V : submodule 𝕜 E} + [complete_space U] (h : U ⟂ V) : + orthogonal_projection U ∘L V.subtypeL = 0 := +continuous_linear_map.ext $ λ v, + orthogonal_projection_mem_subspace_orthogonal_complement_eq_zero $ h.symm v.prop + +/-- The projection into `U` from `V` is the zero map if and only if `U` and `V` are orthogonal. -/ +lemma orthogonal_projection_comp_subtypeL_eq_zero_iff {U V : submodule 𝕜 E} + [complete_space U] : + orthogonal_projection U ∘L V.subtypeL = 0 ↔ U ⟂ V := +⟨λ h u hu v hv, begin + convert orthogonal_projection_inner_eq_zero v u hu using 2, + have : orthogonal_projection U v = 0 := fun_like.congr_fun h ⟨_, hv⟩, + rw [this, submodule.coe_zero, sub_zero] +end, submodule.is_ortho.orthogonal_projection_comp_subtypeL⟩ + lemma orthogonal_projection_eq_linear_proj [complete_space K] (x : E) : orthogonal_projection K x = K.linear_proj_of_is_compl _ submodule.is_compl_orthogonal_of_complete_space x := diff --git a/src/analysis/inner_product_space/spectrum.lean b/src/analysis/inner_product_space/spectrum.lean index f312010c46b87..c40c794b92592 100644 --- a/src/analysis/inner_product_space/spectrum.lean +++ b/src/analysis/inner_product_space/spectrum.lean @@ -107,8 +107,8 @@ lemma orthogonal_supr_eigenspaces (μ : 𝕜) : begin set p : submodule 𝕜 E := (⨆ μ, eigenspace T μ)ᗮ, refine eigenspace_restrict_eq_bot hT.orthogonal_supr_eigenspaces_invariant _, - have H₂ : p ≤ (eigenspace T μ)ᗮ := submodule.orthogonal_le (le_supr _ _), - exact (eigenspace T μ).orthogonal_disjoint.mono_right H₂ + have H₂ : eigenspace T μ ⟂ p := (submodule.is_ortho_orthogonal_right _).mono_left (le_supr _ _), + exact H₂.disjoint end /-! ### Finite-dimensional theory -/ diff --git a/src/geometry/euclidean/monge_point.lean b/src/geometry/euclidean/monge_point.lean index f6eb1d8e00df1..f6e741a38b287 100644 --- a/src/geometry/euclidean/monge_point.lean +++ b/src/geometry/euclidean/monge_point.lean @@ -367,14 +367,11 @@ by rw [altitude_def, /-- The vector span of the opposite face lies in the direction orthogonal to an altitude. -/ -lemma vector_span_le_altitude_direction_orthogonal {n : ℕ} (s : simplex ℝ P (n + 1)) - (i : fin (n + 2)) : - vector_span ℝ (s.points '' ↑(finset.univ.erase i)) ≤ (s.altitude i).directionᗮ := +lemma vector_span_is_ortho_altitude_direction {n : ℕ} (s : simplex ℝ P (n + 1)) (i : fin (n + 2)) : + vector_span ℝ (s.points '' ↑(finset.univ.erase i)) ⟂ (s.altitude i).direction := begin rw direction_altitude, - exact le_trans - (vector_span ℝ (s.points '' ↑(finset.univ.erase i))).le_orthogonal_orthogonal - (submodule.orthogonal_le inf_le_left) + exact (submodule.is_ortho_orthogonal_right _).mono_right inf_le_left, end open finite_dimensional @@ -594,7 +591,7 @@ begin have hle : (t₁.altitude i₃).directionᗮ ≤ line[ℝ, t₁.orthocenter, t₁.points i₃].directionᗮ := submodule.orthogonal_le (direction_le (affine_span_orthocenter_point_le_altitude _ _)), - refine hle ((t₁.vector_span_le_altitude_direction_orthogonal i₃) _), + refine hle ((t₁.vector_span_is_ortho_altitude_direction i₃) _), have hui : finset.univ.erase i₃ = {i₁, i₂}, { clear hle h₂ h₃, dec_trivial! }, rw [hui, finset.coe_insert, finset.coe_singleton, set.image_insert_eq, set.image_singleton], refine vsub_mem_vector_span ℝ (set.mem_insert _ _) From 3b1890e71632be9e3b2086ab512c3259a7e9a3ef Mon Sep 17 00:00:00 2001 From: Alex J Best Date: Tue, 4 Apr 2023 22:00:41 +0000 Subject: [PATCH 0774/1291] chore(topology/algebra/group): generalise instances (#15171) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Using the generalisation linter make the following generalisations in `topology.algebra.group.basic`. Note that topological spaces that are groups with continuous multiplication but noncontinuous inverse do exist at least when the space is noncompact (they are known as paratopological spaces) I do not claim that they are used in mathlib though ;). In summary we generalise: - `tendsto_inv_nhds_within_Ioi` and all variants to only require continuous inverse rather that topological group - the `continuous_inv` operation on the multiplicative opposite to only require a `has_inv`, rather than a group - `topological_group.t1_space` from topological groups to only continuous mul - `topological_group.regular_space` (and `t2_space`) from assuming the base is t1 to just topological group. - `compact_open_separated_mul_right/left`, from topological group to `mul_one_class` with a continuous mul - various `quotient_group.has_continuous_const_smul` type lemmas to continuous_mul - `tsum_sigma`/`tsum_prod` from t1 to t0 and their additivised versions. Co-authored-by: Yaël Dillies Co-authored-by: Eric Wieser --- src/topology/algebra/group/basic.lean | 31 +++++++++++++------- src/topology/algebra/infinite_sum/basic.lean | 16 ++++++---- 2 files changed, 31 insertions(+), 16 deletions(-) diff --git a/src/topology/algebra/group/basic.lean b/src/topology/algebra/group/basic.lean index 4ffba09f58648..a9088e29754f0 100644 --- a/src/topology/algebra/group/basic.lean +++ b/src/topology/algebra/group/basic.lean @@ -414,7 +414,7 @@ end zpow section ordered_comm_group -variables [topological_space H] [ordered_comm_group H] [topological_group H] +variables [topological_space H] [ordered_comm_group H] [has_continuous_inv H] @[to_additive] lemma tendsto_inv_nhds_within_Ioi {a : H} : tendsto has_inv.inv (𝓝[>] a) (𝓝[<] (a⁻¹)) := @@ -463,7 +463,7 @@ instance pi.topological_group {C : β → Type*} [∀ b, topological_space (C b) open mul_opposite @[to_additive] -instance [group α] [has_continuous_inv α] : has_continuous_inv αᵐᵒᵖ := +instance [has_inv α] [has_continuous_inv α] : has_continuous_inv αᵐᵒᵖ := op_homeomorph.symm.inducing.has_continuous_inv unop_inv /-- If multiplication is continuous in `α`, then it also is in `αᵐᵒᵖ`. -/ @@ -1049,12 +1049,17 @@ class add_group_with_zero_nhd (G : Type u) extends add_comm_group G := section filter_mul section -variables (G) [topological_space G] [group G] [topological_group G] +variables (G) [topological_space G] [group G] [has_continuous_mul G] @[to_additive] lemma topological_group.t1_space (h : @is_closed G _ {1}) : t1_space G := ⟨assume x, by { convert is_closed_map_mul_right x _ h, simp }⟩ +end + +section +variables (G) [topological_space G] [group G] [topological_group G] + @[priority 100, to_additive] instance topological_group.regular_space : regular_space G := begin @@ -1071,22 +1076,21 @@ begin end @[to_additive] -lemma topological_group.t3_space [t1_space G] : t3_space G := ⟨⟩ +lemma topological_group.t3_space [t0_space G] : t3_space G := ⟨⟩ @[to_additive] -lemma topological_group.t2_space [t1_space G] : t2_space G := +lemma topological_group.t2_space [t0_space G] : t2_space G := by { haveI := topological_group.t3_space G, apply_instance } variables {G} (S : subgroup G) [subgroup.normal S] [is_closed (S : set G)] @[to_additive] instance subgroup.t3_quotient_of_is_closed - (S : subgroup G) [subgroup.normal S] [is_closed (S : set G)] : t3_space (G ⧸ S) := + (S : subgroup G) [subgroup.normal S] [hS : is_closed (S : set G)] : t3_space (G ⧸ S) := begin - suffices : t1_space (G ⧸ S), { exact @topological_group.t3_space _ _ _ _ this, }, - have hS : is_closed (S : set G) := infer_instance, rw ← quotient_group.ker_mk S at hS, - exact topological_group.t1_space (G ⧸ S) ((quotient_map_quotient_mk.is_closed_preimage).mp hS), + haveI := topological_group.t1_space (G ⧸ S) (quotient_map_quotient_mk.is_closed_preimage.mp hS), + exact topological_group.t3_space _, end /-- A subgroup `S` of a topological group `G` acts on `G` properly discontinuously on the left, if @@ -1141,7 +1145,7 @@ section /-! Some results about an open set containing the product of two sets in a topological group. -/ -variables [topological_space G] [group G] [topological_group G] +variables [topological_space G] [mul_one_class G] [has_continuous_mul G] /-- Given a compact set `K` inside an open set `U`, there is a open neighborhood `V` of `1` such that `K * V ⊆ U`. -/ @@ -1183,6 +1187,11 @@ begin preimage_image_eq _ op_injective] at hV' end +end + +section +variables [topological_space G] [group G] [topological_group G] + /-- A compact set is covered by finitely many left multiplicative translates of a set with non-empty interior. -/ @[to_additive "A compact set is covered by finitely many left additive translates of a set @@ -1293,7 +1302,7 @@ instance {G} [topological_space G] [add_group G] [topological_add_group G] : { continuous_inv := @continuous_neg G _ _ _ } section quotient -variables [group G] [topological_space G] [topological_group G] {Γ : subgroup G} +variables [group G] [topological_space G] [has_continuous_mul G] {Γ : subgroup G} @[to_additive] instance quotient_group.has_continuous_const_smul : has_continuous_const_smul G (G ⧸ Γ) := diff --git a/src/topology/algebra/infinite_sum/basic.lean b/src/topology/algebra/infinite_sum/basic.lean index 1b8274361d364..9d0486d2144d3 100644 --- a/src/topology/algebra/infinite_sum/basic.lean +++ b/src/topology/algebra/infinite_sum/basic.lean @@ -989,8 +989,6 @@ begin exact hde _ (h _ finset.sdiff_disjoint) _ (h _ finset.sdiff_disjoint) } end -local attribute [instance] topological_add_group.t3_space - /-- The sum over the complement of a finset tends to `0` when the finset grows to cover the whole space. This does not need a summability assumption, as otherwise all sums are zero. -/ lemma tendsto_tsum_compl_at_top_zero (f : β → α) : @@ -1075,18 +1073,26 @@ lemma summable.prod_factor {f : β × γ → α} (h : summable f) (b : β) : summable (λ c, f (b, c)) := h.comp_injective $ λ c₁ c₂ h, (prod.ext_iff.1 h).2 -lemma tsum_sigma [t1_space α] {γ : β → Type*} {f : (Σb:β, γ b) → α} +section loc_instances +-- enable inferring a T3-topological space from a topological group +local attribute [instance] topological_add_group.t3_space +-- disable getting a T0-space from a T3-space as this causes loops +local attribute [-instance] t3_space.to_t0_space + +lemma tsum_sigma [t0_space α] {γ : β → Type*} {f : (Σb:β, γ b) → α} (ha : summable f) : ∑'p, f p = ∑'b c, f ⟨b, c⟩ := tsum_sigma' (λ b, ha.sigma_factor b) ha -lemma tsum_prod [t1_space α] {f : β × γ → α} (h : summable f) : +lemma tsum_prod [t0_space α] {f : β × γ → α} (h : summable f) : ∑'p, f p = ∑'b c, f ⟨b, c⟩ := tsum_prod' h h.prod_factor -lemma tsum_comm [t1_space α] {f : β → γ → α} (h : summable (function.uncurry f)) : +lemma tsum_comm [t0_space α] {f : β → γ → α} (h : summable (function.uncurry f)) : ∑' c b, f b c = ∑' b c, f b c := tsum_comm' h h.prod_factor h.prod_symm.prod_factor +end loc_instances + lemma tsum_subtype_add_tsum_subtype_compl [t2_space α] {f : β → α} (hf : summable f) (s : set β) : ∑' x : s, f x + ∑' x : sᶜ, f x = ∑' x, f x := ((hf.subtype s).has_sum.add_compl (hf.subtype {x | x ∉ s}).has_sum).unique hf.has_sum From d64d67d000b974f0d86a2be7918cf800be6271c8 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Wed, 5 Apr 2023 05:31:24 +0000 Subject: [PATCH 0775/1291] chore(*): add mathlib4 synchronization comments (#18726) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebraic_topology.dold_kan.compatibility` * `category_theory.shift.basic` * `category_theory.sites.sheaf_of_types` * `category_theory.triangulated.basic` * `combinatorics.additive.e_transform` * `computability.regular_expressions` * `data.fin.tuple.reflection` * `data.finset.interval` * `data.list.intervals` * `data.matrix.basis` * `data.matrix.block` * `data.matrix.char_p` * `data.matrix.dual_number` * `data.matrix.hadamard` * `data.mv_polynomial.expand` * `data.mv_polynomial.monad` * `linear_algebra.isomorphisms` * `linear_algebra.matrix.orthogonal` * `linear_algebra.matrix.symmetric` * `linear_algebra.matrix.trace` * `logic.hydra` * `ring_theory.finite_type` * `ring_theory.rees_algebra` * `ring_theory.zmod` * `topology.algebra.module.basic` * `topology.uniform_space.matrix` --- src/algebraic_topology/dold_kan/compatibility.lean | 5 ++++- src/category_theory/shift/basic.lean | 3 +++ src/category_theory/sites/sheaf_of_types.lean | 3 +++ src/category_theory/triangulated/basic.lean | 3 +++ src/combinatorics/additive/e_transform.lean | 3 +++ src/computability/regular_expressions.lean | 3 +++ src/data/fin/tuple/reflection.lean | 3 +++ src/data/finset/interval.lean | 3 +++ src/data/list/intervals.lean | 3 +++ src/data/matrix/basis.lean | 3 +++ src/data/matrix/block.lean | 3 +++ src/data/matrix/char_p.lean | 3 +++ src/data/matrix/dual_number.lean | 3 +++ src/data/matrix/hadamard.lean | 3 +++ src/data/mv_polynomial/expand.lean | 3 +++ src/data/mv_polynomial/monad.lean | 3 +++ src/linear_algebra/isomorphisms.lean | 3 +++ src/linear_algebra/matrix/orthogonal.lean | 3 +++ src/linear_algebra/matrix/symmetric.lean | 3 +++ src/linear_algebra/matrix/trace.lean | 3 +++ src/logic/hydra.lean | 3 +++ src/ring_theory/finite_type.lean | 3 +++ src/ring_theory/rees_algebra.lean | 3 +++ src/ring_theory/zmod.lean | 3 +++ src/topology/algebra/module/basic.lean | 3 +++ src/topology/uniform_space/matrix.lean | 3 +++ 26 files changed, 79 insertions(+), 1 deletion(-) diff --git a/src/algebraic_topology/dold_kan/compatibility.lean b/src/algebraic_topology/dold_kan/compatibility.lean index 8093d7bfa79ce..b5afeedb24a51 100644 --- a/src/algebraic_topology/dold_kan/compatibility.lean +++ b/src/algebraic_topology/dold_kan/compatibility.lean @@ -6,7 +6,10 @@ Authors: Joël Riou import category_theory.equivalence -/-! Tools for compatibilities between Dold-Kan equivalences +/-! +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Tools for compatibilities between Dold-Kan equivalences The purpose of this file is to introduce tools which will enable the construction of the Dold-Kan equivalence `simplicial_object C ≌ chain_complex C ℕ` diff --git a/src/category_theory/shift/basic.lean b/src/category_theory/shift/basic.lean index d78bf92733d19..fb22baa97b3b6 100644 --- a/src/category_theory/shift/basic.lean +++ b/src/category_theory/shift/basic.lean @@ -10,6 +10,9 @@ import category_theory.monoidal.discrete /-! # Shift +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A `shift` on a category `C` indexed by a monoid `A` is nothing more than a monoidal functor from `A` to `C ⥤ C`. A typical example to keep in mind might be the category of complexes `⋯ → C_{n-1} → C_n → C_{n+1} → ⋯`. It has a shift indexed by `ℤ`, where we assign to diff --git a/src/category_theory/sites/sheaf_of_types.lean b/src/category_theory/sites/sheaf_of_types.lean index 55220cad13759..cdd033564ed27 100644 --- a/src/category_theory/sites/sheaf_of_types.lean +++ b/src/category_theory/sites/sheaf_of_types.lean @@ -10,6 +10,9 @@ import category_theory.limits.shapes.types /-! # Sheaves of types on a Grothendieck topology +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Defines the notion of a sheaf of types (usually called a sheaf of sets by mathematicians) on a category equipped with a Grothendieck topology, as well as a range of equivalent conditions useful in different situations. diff --git a/src/category_theory/triangulated/basic.lean b/src/category_theory/triangulated/basic.lean index 8b1ab1d9950bf..94740f2468fd0 100644 --- a/src/category_theory/triangulated/basic.lean +++ b/src/category_theory/triangulated/basic.lean @@ -9,6 +9,9 @@ import category_theory.shift.basic /-! # Triangles +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains the definition of triangles in an additive category with an additive shift. It also defines morphisms between these triangles. diff --git a/src/combinatorics/additive/e_transform.lean b/src/combinatorics/additive/e_transform.lean index 1f99c723f6a8a..cee3e26fab0fd 100644 --- a/src/combinatorics/additive/e_transform.lean +++ b/src/combinatorics/additive/e_transform.lean @@ -8,6 +8,9 @@ import data.finset.pointwise /-! # e-transforms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + e-transforms are a family of transformations of pairs of finite sets that aim to reduce the size of the sumset while keeping some invariant the same. This file defines a few of them, to be used as internals of other proofs. diff --git a/src/computability/regular_expressions.lean b/src/computability/regular_expressions.lean index d1f36e987068e..d1043d9886779 100644 --- a/src/computability/regular_expressions.lean +++ b/src/computability/regular_expressions.lean @@ -9,6 +9,9 @@ import computability.language /-! # Regular Expressions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains the formal definition for regular expressions and basic lemmas. Note these are regular expressions in terms of formal language theory. Note this is different to regex's used in computer science such as the POSIX standard. diff --git a/src/data/fin/tuple/reflection.lean b/src/data/fin/tuple/reflection.lean index 413c3ef0f4ebe..9c338280ad8c5 100644 --- a/src/data/fin/tuple/reflection.lean +++ b/src/data/fin/tuple/reflection.lean @@ -9,6 +9,9 @@ import algebra.big_operators.fin /-! # Lemmas for tuples `fin m → α` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains alternative definitions of common operators on vectors which expand definitionally to the expected expression when evaluated on `![]` notation. diff --git a/src/data/finset/interval.lean b/src/data/finset/interval.lean index 3f1181ef49391..a8e5585b4c7da 100644 --- a/src/data/finset/interval.lean +++ b/src/data/finset/interval.lean @@ -8,6 +8,9 @@ import data.finset.locally_finite /-! # Intervals of finsets as finsets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file provides the `locally_finite_order` instance for `finset α` and calculates the cardinality of finite intervals of finsets. diff --git a/src/data/list/intervals.lean b/src/data/list/intervals.lean index 0ece01cfaa3ef..0e6984b4b3ccd 100644 --- a/src/data/list/intervals.lean +++ b/src/data/list/intervals.lean @@ -9,6 +9,9 @@ import data.list.range /-! # Intervals in ℕ +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines intervals of naturals. `list.Ico m n` is the list of integers greater than `m` and strictly less than `n`. diff --git a/src/data/matrix/basis.lean b/src/data/matrix/basis.lean index cb40e55b8f0b4..276ed52bb3a49 100644 --- a/src/data/matrix/basis.lean +++ b/src/data/matrix/basis.lean @@ -9,6 +9,9 @@ import linear_algebra.matrix.trace /-! # Matrices with a single non-zero element. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file provides `matrix.std_basis_matrix`. The matrix `matrix.std_basis_matrix i j c` has `c` at position `(i, j)`, and zeroes elsewhere. -/ diff --git a/src/data/matrix/block.lean b/src/data/matrix/block.lean index d51d3e867abc4..2b48884753454 100644 --- a/src/data/matrix/block.lean +++ b/src/data/matrix/block.lean @@ -8,6 +8,9 @@ import data.matrix.basic /-! # Block Matrices +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions * `matrix.from_blocks`: build a block matrix out of 4 blocks diff --git a/src/data/matrix/char_p.lean b/src/data/matrix/char_p.lean index 49e8e4e21776e..85e4b06549949 100644 --- a/src/data/matrix/char_p.lean +++ b/src/data/matrix/char_p.lean @@ -7,6 +7,9 @@ import data.matrix.basic import algebra.char_p.basic /-! # Matrices in prime characteristic + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ open matrix diff --git a/src/data/matrix/dual_number.lean b/src/data/matrix/dual_number.lean index 256458b40e69e..3f9e9c3ad10e4 100644 --- a/src/data/matrix/dual_number.lean +++ b/src/data/matrix/dual_number.lean @@ -9,6 +9,9 @@ import data.matrix.basic /-! # Matrices of dual numbers are isomorphic to dual numbers over matrices +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Showing this for the more general case of `triv_sq_zero_ext R M` would require an action between `matrix n n R` and `matrix n n M`, which would risk causing diamonds. -/ diff --git a/src/data/matrix/hadamard.lean b/src/data/matrix/hadamard.lean index a95155afaf89e..d8ca3ffc18d91 100644 --- a/src/data/matrix/hadamard.lean +++ b/src/data/matrix/hadamard.lean @@ -8,6 +8,9 @@ import linear_algebra.matrix.trace /-! # Hadamard product of matrices +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the Hadamard product `matrix.hadamard` and contains basic properties about them. diff --git a/src/data/mv_polynomial/expand.lean b/src/data/mv_polynomial/expand.lean index ab49cf584889e..13a0bac69cb50 100644 --- a/src/data/mv_polynomial/expand.lean +++ b/src/data/mv_polynomial/expand.lean @@ -8,6 +8,9 @@ import data.mv_polynomial.monad /-! ## Expand multivariate polynomials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given a multivariate polynomial `φ`, one may replace every occurence of `X i` by `X i ^ n`, for some natural number `n`. This operation is called `mv_polynomial.expand` and it is an algebra homomorphism. diff --git a/src/data/mv_polynomial/monad.lean b/src/data/mv_polynomial/monad.lean index a59fb34f870b7..106cefe8a1c1d 100644 --- a/src/data/mv_polynomial/monad.lean +++ b/src/data/mv_polynomial/monad.lean @@ -10,6 +10,9 @@ import data.mv_polynomial.variables # Monad operations on `mv_polynomial` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines two monadic operations on `mv_polynomial`. Given `p : mv_polynomial σ R`, * `mv_polynomial.bind₁` and `mv_polynomial.join₁` operate on the variable type `σ`. diff --git a/src/linear_algebra/isomorphisms.lean b/src/linear_algebra/isomorphisms.lean index b77b3abdc038b..94799845bdfa5 100644 --- a/src/linear_algebra/isomorphisms.lean +++ b/src/linear_algebra/isomorphisms.lean @@ -8,6 +8,9 @@ import linear_algebra.quotient /-! # Isomorphism theorems for modules. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + * The Noether's first, second, and third isomorphism theorems for modules are proved as `linear_map.quot_ker_equiv_range`, `linear_map.quotient_inf_equiv_sup_quotient` and `submodule.quotient_quotient_equiv_quotient`. diff --git a/src/linear_algebra/matrix/orthogonal.lean b/src/linear_algebra/matrix/orthogonal.lean index 8ac01b340accc..561aa1d1c46ae 100644 --- a/src/linear_algebra/matrix/orthogonal.lean +++ b/src/linear_algebra/matrix/orthogonal.lean @@ -8,6 +8,9 @@ import data.matrix.basic /-! # Orthogonal +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains definitions and properties concerning orthogonality of rows and columns. ## Main results diff --git a/src/linear_algebra/matrix/symmetric.lean b/src/linear_algebra/matrix/symmetric.lean index 5b1c8525d673d..108fb0f1d3e67 100644 --- a/src/linear_algebra/matrix/symmetric.lean +++ b/src/linear_algebra/matrix/symmetric.lean @@ -8,6 +8,9 @@ import data.matrix.block /-! # Symmetric matrices +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains the definition and basic results about symmetric matrices. ## Main definition diff --git a/src/linear_algebra/matrix/trace.lean b/src/linear_algebra/matrix/trace.lean index 9b642f3f2c175..b46f3b834f19b 100644 --- a/src/linear_algebra/matrix/trace.lean +++ b/src/linear_algebra/matrix/trace.lean @@ -8,6 +8,9 @@ import data.matrix.basic /-! # Trace of a matrix +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the trace of a matrix, the map sending a matrix to the sum of its diagonal entries. diff --git a/src/logic/hydra.lean b/src/logic/hydra.lean index 7a6c92e116e0f..97260914248eb 100644 --- a/src/logic/hydra.lean +++ b/src/logic/hydra.lean @@ -10,6 +10,9 @@ import order.game_add /-! # Termination of a hydra game +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file deals with the following version of the hydra game: each head of the hydra is labelled by an element in a type `α`, and when you cut off one head with label `a`, it grows back an arbitrary but finite number of heads, all labelled by elements smaller than diff --git a/src/ring_theory/finite_type.lean b/src/ring_theory/finite_type.lean index b862b84dff4d2..b0c17406ed9b7 100644 --- a/src/ring_theory/finite_type.lean +++ b/src/ring_theory/finite_type.lean @@ -12,6 +12,9 @@ import ring_theory.noetherian /-! # Finiteness conditions in commutative algebra +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define a notion of finiteness that is common in commutative algebra. ## Main declarations diff --git a/src/ring_theory/rees_algebra.lean b/src/ring_theory/rees_algebra.lean index 437facd972ede..8668d8794b884 100644 --- a/src/ring_theory/rees_algebra.lean +++ b/src/ring_theory/rees_algebra.lean @@ -10,6 +10,9 @@ import ring_theory.finite_type # Rees algebra +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The Rees algebra of an ideal `I` is the subalgebra `R[It]` of `R[t]` defined as `R[It] = ⨁ₙ Iⁿ tⁿ`. This is used to prove the Artin-Rees lemma, and will potentially enable us to calculate some blowup in the future. diff --git a/src/ring_theory/zmod.lean b/src/ring_theory/zmod.lean index 62b6858bc9436..85aa3d4eed097 100644 --- a/src/ring_theory/zmod.lean +++ b/src/ring_theory/zmod.lean @@ -10,6 +10,9 @@ import ring_theory.int.basic /-! # Ring theoretic facts about `zmod n` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We collect a few facts about `zmod n` that need some ring theory to be proved/stated ## Main statements diff --git a/src/topology/algebra/module/basic.lean b/src/topology/algebra/module/basic.lean index 6d322735e4f11..aafd5fefb87f1 100644 --- a/src/topology/algebra/module/basic.lean +++ b/src/topology/algebra/module/basic.lean @@ -16,6 +16,9 @@ import linear_algebra.pi /-! # Theory of topological modules and continuous linear maps. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We use the class `has_continuous_smul` for topological (semi) modules and topological vector spaces. In this file we define continuous (semi-)linear maps, as semilinear maps between topological diff --git a/src/topology/uniform_space/matrix.lean b/src/topology/uniform_space/matrix.lean index 2ad9ecbce6da8..476be2d9ca9d3 100644 --- a/src/topology/uniform_space/matrix.lean +++ b/src/topology/uniform_space/matrix.lean @@ -8,6 +8,9 @@ import data.matrix.basic /-! # Uniform space structure on matrices + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ open_locale uniformity topology From e8638a0fcaf73e4500469f368ef9494e495099b3 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 5 Apr 2023 07:48:13 +0000 Subject: [PATCH 0776/1291] feat(algebra/divisibility/basic): Dot notation aliases (#18698) A few convenience shortcuts for `dvd` along with some simple `nat` lemmas. Also * Drop `neg_dvd_of_dvd`/`dvd_of_neg_dvd`/`dvd_neg_of_dvd`/`dvd_of_dvd_neg` in favor of the aforementioned shortcuts. * Remove explicit arguments to `dvd_neg`/`neg_dvd`. * Drop `int.of_nat_dvd_of_dvd_nat_abs`/`int.dvd_nat_abs_of_of_nat_dvd` because they are the two directions of `int.coe_nat_dvd_left`. * Move `group_with_zero.to_cancel_monoid_with_zero` from `algebra.group_with_zero.units.basic` back to `algebra.group_with_zero.basic`. It was erroneously moved during the Great Splits. --- archive/imo/imo1988_q6.lean | 2 +- archive/imo/imo2011_q5.lean | 2 +- src/algebra/divisibility/basic.lean | 6 +- src/algebra/euclidean_domain/basic.lean | 2 +- src/algebra/group/units.lean | 11 +++ src/algebra/group_with_zero/basic.lean | 17 ++++ src/algebra/group_with_zero/divisibility.lean | 25 ++++++ src/algebra/order/monoid/canonical/defs.lean | 2 + src/algebra/ring/boolean_ring.lean | 4 +- src/algebra/ring/divisibility.lean | 81 +++++++++---------- src/data/int/dvd/basic.lean | 16 ---- src/data/int/dvd/pow.lean | 17 +--- src/data/int/order/basic.lean | 2 +- src/data/nat/gcd/basic.lean | 3 + src/data/nat/order/basic.lean | 14 +--- src/data/nat/order/lemmas.lean | 1 + src/group_theory/perm/cycle/basic.lean | 2 +- src/number_theory/arithmetic_function.lean | 2 +- src/number_theory/divisors.lean | 2 +- src/number_theory/multiplicity.lean | 2 +- src/number_theory/pythagorean_triples.lean | 7 +- src/number_theory/zsqrtd/basic.lean | 2 +- src/number_theory/zsqrtd/gaussian_int.lean | 4 +- src/ring_theory/multiplicity.lean | 8 +- src/tactic/omega/eq_elim.lean | 4 +- 25 files changed, 128 insertions(+), 110 deletions(-) diff --git a/archive/imo/imo1988_q6.lean b/archive/imo/imo1988_q6.lean index 24f0137c54364..044358ba8f7bd 100644 --- a/archive/imo/imo1988_q6.lean +++ b/archive/imo/imo1988_q6.lean @@ -266,7 +266,7 @@ begin have x_sq_dvd : x*x ∣ x*x*k := dvd_mul_right (x*x) k, rw ← hx at x_sq_dvd, obtain ⟨y, hy⟩ : x * x ∣ 1 := by simpa only [nat.dvd_add_self_left, add_assoc] using x_sq_dvd, - obtain ⟨rfl,rfl⟩ : x = 1 ∧ y = 1 := by simpa [nat.mul_eq_one_iff] using hy.symm, + obtain ⟨rfl,rfl⟩ : x = 1 ∧ y = 1 := by simpa [mul_eq_one] using hy.symm, simpa using hx.symm, }, { -- Show the descent step. intros x y x_lt_y hx h_base h z h_root hV₁ hV₀, diff --git a/archive/imo/imo2011_q5.lean b/archive/imo/imo2011_q5.lean index a386587c31a31..abf6897ad352d 100644 --- a/archive/imo/imo2011_q5.lean +++ b/archive/imo/imo2011_q5.lean @@ -47,7 +47,7 @@ begin { -- d = 0 exact hd }, { -- d < 0 - have h₁ : f n ≤ -d, from le_of_dvd (neg_pos.mpr hd) ((dvd_neg _ _).mpr h_fn_dvd_d), + have h₁ : f n ≤ -d, from le_of_dvd (neg_pos.mpr hd) h_fn_dvd_d.neg_right, have h₂ : ¬ f n ≤ -d, from not_le.mpr h_neg_d_lt_fn, contradiction } }, have h₁ : f m = f (m - n), from sub_eq_zero.mp h_d_eq_zero, diff --git a/src/algebra/divisibility/basic.lean b/src/algebra/divisibility/basic.lean index ffa10bb7ff9c1..3495d550e3400 100644 --- a/src/algebra/divisibility/basic.lean +++ b/src/algebra/divisibility/basic.lean @@ -91,7 +91,7 @@ end semigroup section monoid -variables [monoid α] +variables [monoid α] {a b : α} @[refl, simp] theorem dvd_refl (a : α) : a ∣ a := dvd.intro 1 (mul_one a) theorem dvd_rfl : ∀ {a : α}, a ∣ a := dvd_refl @@ -99,6 +99,10 @@ instance : is_refl α (∣) := ⟨dvd_refl⟩ theorem one_dvd (a : α) : 1 ∣ a := dvd.intro a (one_mul a) +lemma dvd_of_eq (h : a = b) : a ∣ b := by rw h + +alias dvd_of_eq ← eq.dvd + end monoid section comm_semigroup diff --git a/src/algebra/euclidean_domain/basic.lean b/src/algebra/euclidean_domain/basic.lean index 3f485ba27a7f9..fde4f36389841 100644 --- a/src/algebra/euclidean_domain/basic.lean +++ b/src/algebra/euclidean_domain/basic.lean @@ -54,7 +54,7 @@ by { rw mul_comm, exact mul_div_cancel_left a b0 } mod_eq_zero.2 dvd_rfl lemma dvd_mod_iff {a b c : R} (h : c ∣ b) : c ∣ a % b ↔ c ∣ a := -by rw [dvd_add_iff_right (h.mul_right _), div_add_mod] +by rw [←dvd_add_right (h.mul_right _), div_add_mod] @[simp] lemma mod_one (a : R) : a % 1 = 0 := mod_eq_zero.2 (one_dvd _) diff --git a/src/algebra/group/units.lean b/src/algebra/group/units.lean index d22ea64c3f1d4..a8d5ad8fdc841 100644 --- a/src/algebra/group/units.lean +++ b/src/algebra/group/units.lean @@ -328,6 +328,17 @@ by rw [divp_eq_iff_mul_eq, divp_mul_eq_mul_divp, divp_eq_iff_mul_eq] (x /ₚ ux) * (y /ₚ uy) = (x * y) /ₚ (ux * uy) := by rw [divp_mul_eq_mul_divp, divp_assoc', divp_divp_eq_divp_mul] +variables [subsingleton αˣ] {a b : α} + +@[to_additive] lemma eq_one_of_mul_right (h : a * b = 1) : a = 1 := +congr_arg units.inv $ subsingleton.elim (units.mk _ _ (by rwa mul_comm) h) 1 + +@[to_additive] lemma eq_one_of_mul_left (h : a * b = 1) : b = 1 := +congr_arg units.inv $ subsingleton.elim (units.mk _ _ h $ by rwa mul_comm) 1 + +@[simp, to_additive] lemma mul_eq_one : a * b = 1 ↔ a = 1 ∧ b = 1 := +⟨λ h, ⟨eq_one_of_mul_right h, eq_one_of_mul_left h⟩, by { rintro ⟨rfl, rfl⟩, exact mul_one _ }⟩ + end comm_monoid /-! diff --git a/src/algebra/group_with_zero/basic.lean b/src/algebra/group_with_zero/basic.lean index 457883ae90ddd..ab120f32d347f 100644 --- a/src/algebra/group_with_zero/basic.lean +++ b/src/algebra/group_with_zero/basic.lean @@ -156,6 +156,15 @@ lemma mul_left_eq_self₀ : a * b = b ↔ a = 1 ∨ b = 0 := calc a * b = b ↔ a * b = 1 * b : by rw one_mul ... ↔ a = 1 ∨ b = 0 : mul_eq_mul_right_iff +@[simp] lemma mul_eq_left₀ (ha : a ≠ 0) : a * b = a ↔ b = 1 := +by rw [iff.comm, ←mul_right_inj' ha, mul_one] + +@[simp] lemma mul_eq_right₀ (hb : b ≠ 0) : a * b = b ↔ a = 1 := +by rw [iff.comm, ←mul_left_inj' hb, one_mul] + +@[simp] lemma left_eq_mul₀ (ha : a ≠ 0) : a = a * b ↔ b = 1 := by rw [eq_comm, mul_eq_left₀ ha] +@[simp] lemma right_eq_mul₀ (hb : b ≠ 0) : b = a * b ↔ a = 1 := by rw [eq_comm, mul_eq_right₀ hb] + /-- An element of a `cancel_monoid_with_zero` fixed by right multiplication by an element other than one must be zero. -/ theorem eq_zero_of_mul_eq_self_right (h₁ : b ≠ 1) (h₂ : a * b = a) : a = 0 := @@ -228,6 +237,14 @@ instance group_with_zero.to_division_monoid : division_monoid G₀ := inv_eq_of_mul := λ a b, inv_eq_of_mul, ..‹group_with_zero G₀› } +@[priority 10] -- see Note [lower instance priority] +instance group_with_zero.to_cancel_monoid_with_zero : cancel_monoid_with_zero G₀ := +{ mul_left_cancel_of_ne_zero := λ x y z hx h, + by rw [← inv_mul_cancel_left₀ hx y, h, inv_mul_cancel_left₀ hx z], + mul_right_cancel_of_ne_zero := λ x y z hy h, + by rw [← mul_inv_cancel_right₀ hy x, h, mul_inv_cancel_right₀ hy z], + ..‹group_with_zero G₀› } + end group_with_zero section group_with_zero diff --git a/src/algebra/group_with_zero/divisibility.lean b/src/algebra/group_with_zero/divisibility.lean index ae8f2bc07b7ad..47ad0aff0b759 100644 --- a/src/algebra/group_with_zero/divisibility.lean +++ b/src/algebra/group_with_zero/divisibility.lean @@ -88,3 +88,28 @@ begin end end monoid_with_zero + +section cancel_comm_monoid_with_zero +variables [cancel_comm_monoid_with_zero α] [subsingleton αˣ] {a b : α} + +lemma dvd_antisymm : a ∣ b → b ∣ a → a = b := +begin + rintro ⟨c, rfl⟩ ⟨d, hcd⟩, + rw [mul_assoc, eq_comm, mul_right_eq_self₀, mul_eq_one] at hcd, + obtain ⟨rfl, -⟩ | rfl := hcd; simp, +end + +attribute [protected] nat.dvd_antisymm --This lemma is in core, so we protect it here + +lemma dvd_antisymm' : a ∣ b → b ∣ a → b = a := flip dvd_antisymm + +alias dvd_antisymm ← has_dvd.dvd.antisymm +alias dvd_antisymm' ← has_dvd.dvd.antisymm' + +lemma eq_of_forall_dvd (h : ∀ c, a ∣ c ↔ b ∣ c) : a = b := +((h _).2 dvd_rfl).antisymm $ (h _).1 dvd_rfl + +lemma eq_of_forall_dvd' (h : ∀ c, c ∣ a ↔ c ∣ b) : a = b := +((h _).1 dvd_rfl).antisymm $ (h _).2 dvd_rfl + +end cancel_comm_monoid_with_zero diff --git a/src/algebra/order/monoid/canonical/defs.lean b/src/algebra/order/monoid/canonical/defs.lean index 03ca16ab30062..ddce37a99ab51 100644 --- a/src/algebra/order/monoid/canonical/defs.lean +++ b/src/algebra/order/monoid/canonical/defs.lean @@ -140,6 +140,8 @@ le_iff_exists_mul.mpr ⟨a, (one_mul _).symm⟩ @[to_additive] lemma bot_eq_one : (⊥ : α) = 1 := le_antisymm bot_le (one_le ⊥) +--TODO: This is a special case of `mul_eq_one`. We need the instance +-- `canonically_ordered_monoid α → unique αˣ` @[simp, to_additive] lemma mul_eq_one_iff : a * b = 1 ↔ a = 1 ∧ b = 1 := mul_eq_one_iff' (one_le _) (one_le _) diff --git a/src/algebra/ring/boolean_ring.lean b/src/algebra/ring/boolean_ring.lean index 53e5f8de7853f..f93c74cd170fc 100644 --- a/src/algebra/ring/boolean_ring.lean +++ b/src/algebra/ring/boolean_ring.lean @@ -63,7 +63,7 @@ calc -a = -a + 0 : by rw add_zero ... = -a + -a + a : by rw [←neg_add_self, add_assoc] ... = a : by rw [add_self, zero_add] -lemma add_eq_zero : a + b = 0 ↔ a = b := +lemma add_eq_zero' : a + b = 0 ↔ a = b := calc a + b = 0 ↔ a = -b : add_eq_zero_iff_eq_neg ... ↔ a = b : by rw neg_eq @@ -82,7 +82,7 @@ by rw [sub_eq_add_neg, add_right_inj, neg_eq] @[priority 100] -- Note [lower instance priority] instance boolean_ring.to_comm_ring : comm_ring α := -{ mul_comm := λ a b, by rw [←add_eq_zero, mul_add_mul], +{ mul_comm := λ a b, by rw [←add_eq_zero', mul_add_mul], .. (infer_instance : boolean_ring α) } end boolean_ring diff --git a/src/algebra/ring/divisibility.lean b/src/algebra/ring/divisibility.lean index 00ee02cd96226..2a44107e6500f 100644 --- a/src/algebra/ring/divisibility.lean +++ b/src/algebra/ring/divisibility.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Leonardo de Moura, Floris van Doorn, Yury Kudryashov, Neil Strickland -/ import algebra.divisibility.basic +import algebra.hom.equiv.basic import algebra.ring.defs /-! @@ -21,6 +22,8 @@ variables [has_add α] [semigroup α] theorem dvd_add [left_distrib_class α] {a b c : α} (h₁ : a ∣ b) (h₂ : a ∣ c) : a ∣ b + c := dvd.elim h₁ (λ d hd, dvd.elim h₂ (λ e he, dvd.intro (d + e) (by simp [left_distrib, hd, he]))) +alias dvd_add ← has_dvd.dvd.add + end distrib_semigroup @[simp] theorem two_dvd_bit0 [semiring α] {a : α} : 2 ∣ bit0 a := ⟨a, bit0_eq_two_mul _⟩ @@ -39,27 +42,16 @@ section semigroup variables [semigroup α] [has_distrib_neg α] {a b c : α} -theorem dvd_neg_of_dvd (h : a ∣ b) : (a ∣ -b) := -let ⟨c, hc⟩ := h in ⟨-c, by simp [hc]⟩ - -theorem dvd_of_dvd_neg (h : a ∣ -b) : (a ∣ b) := -let t := dvd_neg_of_dvd h in by rwa neg_neg at t - -/-- An element a of a semigroup with a distributive negation divides the negation of an element b -iff a divides b. -/ -@[simp] lemma dvd_neg (a b : α) : (a ∣ -b) ↔ (a ∣ b) := -⟨dvd_of_dvd_neg, dvd_neg_of_dvd⟩ +/-- An element `a` of a semigroup with a distributive negation divides the negation of an element +`b` iff `a` divides `b`. -/ +@[simp] lemma dvd_neg : a ∣ -b ↔ a ∣ b := (equiv.neg _).exists_congr_left.trans $ by simpa -theorem neg_dvd_of_dvd (h : a ∣ b) : -a ∣ b := -let ⟨c, hc⟩ := h in ⟨-c, by simp [hc]⟩ +/-- The negation of an element `a` of a semigroup with a distributive negation divides another +element `b` iff `a` divides `b`. -/ +@[simp] lemma neg_dvd : -a ∣ b ↔ a ∣ b := (equiv.neg _).exists_congr_left.trans $ by simpa -theorem dvd_of_neg_dvd (h : -a ∣ b) : a ∣ b := -let t := neg_dvd_of_dvd h in by rwa neg_neg at t - -/-- The negation of an element a of a semigroup with a distributive negation divides -another element b iff a divides b. -/ -@[simp] lemma neg_dvd (a b : α) : (-a ∣ b) ↔ (a ∣ b) := -⟨dvd_of_neg_dvd, neg_dvd_of_dvd⟩ +alias neg_dvd ↔ has_dvd.dvd.of_neg_left has_dvd.dvd.neg_left +alias dvd_neg ↔ has_dvd.dvd.of_neg_right has_dvd.dvd.neg_right end semigroup @@ -67,34 +59,31 @@ section non_unital_ring variables [non_unital_ring α] {a b c : α} theorem dvd_sub (h₁ : a ∣ b) (h₂ : a ∣ c) : a ∣ b - c := -by { rw sub_eq_add_neg, exact dvd_add h₁ (dvd_neg_of_dvd h₂) } - -theorem dvd_add_iff_left (h : a ∣ c) : a ∣ b ↔ a ∣ b + c := -⟨λh₂, dvd_add h₂ h, λH, by have t := dvd_sub H h; rwa add_sub_cancel at t⟩ +by simpa only [←sub_eq_add_neg] using h₁.add h₂.neg_right -theorem dvd_add_iff_right (h : a ∣ b) : a ∣ c ↔ a ∣ b + c := -by rw add_comm; exact dvd_add_iff_left h +alias dvd_sub ← has_dvd.dvd.sub -/-- If an element a divides another element c in a commutative ring, a divides the sum of another - element b with c iff a divides b. -/ +/-- If an element `a` divides another element `c` in a ring, `a` divides the sum of another element +`b` with `c` iff `a` divides `b`. -/ theorem dvd_add_left (h : a ∣ c) : a ∣ b + c ↔ a ∣ b := -(dvd_add_iff_left h).symm +⟨λ H, by simpa only [add_sub_cancel] using dvd_sub H h, λ h₂, dvd_add h₂ h⟩ -/-- If an element a divides another element b in a commutative ring, a divides the sum of b and - another element c iff a divides c. -/ -theorem dvd_add_right (h : a ∣ b) : a ∣ b + c ↔ a ∣ c := -(dvd_add_iff_right h).symm +/-- If an element `a` divides another element `b` in a ring, `a` divides the sum of `b` and another +element `c` iff `a` divides `c`. -/ +theorem dvd_add_right (h : a ∣ b) : a ∣ b + c ↔ a ∣ c := by rw add_comm; exact dvd_add_left h -lemma dvd_iff_dvd_of_dvd_sub {a b c : α} (h : a ∣ (b - c)) : (a ∣ b ↔ a ∣ c) := -begin - split, - { intro h', - convert dvd_sub h' h, - exact eq.symm (sub_sub_self b c) }, - { intro h', - convert dvd_add h h', - exact eq_add_of_sub_eq rfl } -end +/-- If an element `a` divides another element `c` in a ring, `a` divides the difference of another +element `b` with `c` iff `a` divides `b`. -/ +theorem dvd_sub_left (h : a ∣ c) : a ∣ b - c ↔ a ∣ b := +by simpa only [←sub_eq_add_neg] using dvd_add_left (dvd_neg.2 h) + +/-- If an element `a` divides another element `b` in a ring, `a` divides the difference of `b` and +another element `c` iff `a` divides `c`. -/ +theorem dvd_sub_right (h : a ∣ b) : a ∣ b - c ↔ a ∣ c := +by rw [sub_eq_add_neg, dvd_add_right h, dvd_neg] + +lemma dvd_iff_dvd_of_dvd_sub (h : a ∣ b - c) : a ∣ b ↔ a ∣ c := +by rw [←sub_add_cancel b c, dvd_add_right h] lemma dvd_sub_comm : a ∣ b - c ↔ a ∣ c - b := by rw [←dvd_neg, neg_sub] @@ -103,7 +92,7 @@ end non_unital_ring section ring variables [ring α] {a b c : α} -theorem two_dvd_bit1 : 2 ∣ bit1 a ↔ (2 : α) ∣ 1 := (dvd_add_iff_right (@two_dvd_bit0 _ _ a)).symm +theorem two_dvd_bit1 : 2 ∣ bit1 a ↔ (2 : α) ∣ 1 := dvd_add_right two_dvd_bit0 /-- An element a divides the sum a + b if and only if a divides b.-/ @[simp] lemma dvd_add_self_left {a b : α} : a ∣ a + b ↔ a ∣ b := @@ -113,6 +102,12 @@ dvd_add_right (dvd_refl a) @[simp] lemma dvd_add_self_right {a b : α} : a ∣ b + a ↔ a ∣ b := dvd_add_left (dvd_refl a) +/-- An element `a` divides the difference `a - b` if and only if `a` divides `b`. -/ +@[simp] lemma dvd_sub_self_left : a ∣ a - b ↔ a ∣ b := dvd_sub_right dvd_rfl + +/-- An element `a` divides the difference `b - a` if and only if `a` divides `b`. -/ +@[simp] lemma dvd_sub_self_right : a ∣ b - a ↔ a ∣ b := dvd_sub_left dvd_rfl + end ring section non_unital_comm_ring diff --git a/src/data/int/dvd/basic.lean b/src/data/int/dvd/basic.lean index 70b50fe2e31a9..8b1da123a6620 100644 --- a/src/data/int/dvd/basic.lean +++ b/src/data/int/dvd/basic.lean @@ -52,22 +52,6 @@ eq_one_of_dvd_one H ⟨b, H'.symm⟩ theorem eq_one_of_mul_eq_one_left {a b : ℤ} (H : 0 ≤ b) (H' : a * b = 1) : b = 1 := eq_one_of_mul_eq_one_right H (by rw [mul_comm, H']) -lemma of_nat_dvd_of_dvd_nat_abs {a : ℕ} : ∀ {z : ℤ} (haz : a ∣ z.nat_abs), ↑a ∣ z -| (int.of_nat _) haz := int.coe_nat_dvd.2 haz -| -[1+k] haz := - begin - change ↑a ∣ -(k+1 : ℤ), - apply dvd_neg_of_dvd, - apply int.coe_nat_dvd.2, - exact haz - end - -lemma dvd_nat_abs_of_of_nat_dvd {a : ℕ} : ∀ {z : ℤ} (haz : ↑a ∣ z), a ∣ z.nat_abs -| (int.of_nat _) haz := int.coe_nat_dvd.1 (int.dvd_nat_abs.2 haz) -| -[1+k] haz := - have haz' : (↑a:ℤ) ∣ (↑(k+1):ℤ), from dvd_of_dvd_neg haz, - int.coe_nat_dvd.1 haz' - theorem dvd_antisymm {a b : ℤ} (H1 : 0 ≤ a) (H2 : 0 ≤ b) : a ∣ b → b ∣ a → a = b := begin rw [← abs_of_nonneg H1, ← abs_of_nonneg H2, abs_eq_nat_abs, abs_eq_nat_abs], diff --git a/src/data/int/dvd/pow.lean b/src/data/int/dvd/pow.lean index 1c74d4e8d51a7..aab19316e3ea5 100644 --- a/src/data/int/dvd/pow.lean +++ b/src/data/int/dvd/pow.lean @@ -23,23 +23,12 @@ theorem sign_pow_bit1 (k : ℕ) : ∀ n : ℤ, n.sign ^ (bit1 k) = n.sign | 0 := zero_pow (nat.zero_lt_bit1 k) | -[1+ n] := (neg_pow_bit1 1 k).trans (congr_arg (λ x, -x) (one_pow (bit1 k))) +--TODO: Do we really need this lemma? lemma pow_dvd_of_le_of_pow_dvd {p m n : ℕ} {k : ℤ} (hmn : m ≤ n) (hdiv : ↑(p ^ n) ∣ k) : ↑(p ^ m) ∣ k := -begin - induction k, - { apply int.coe_nat_dvd.2, - apply pow_dvd_of_le_of_pow_dvd hmn, - apply int.coe_nat_dvd.1 hdiv }, - change -[1+k] with -(↑(k+1) : ℤ), - apply dvd_neg_of_dvd, - apply int.coe_nat_dvd.2, - apply pow_dvd_of_le_of_pow_dvd hmn, - apply int.coe_nat_dvd.1, - apply dvd_of_dvd_neg, - exact hdiv, -end +(pow_dvd_pow _ hmn).nat_cast.trans hdiv lemma dvd_of_pow_dvd {p k : ℕ} {m : ℤ} (hk : 1 ≤ k) (hpk : ↑(p^k) ∣ m) : ↑p ∣ m := -by rw ←pow_one p; exact pow_dvd_of_le_of_pow_dvd hk hpk +(dvd_pow_self _ $ pos_iff_ne_zero.1 hk).nat_cast.trans hpk end int diff --git a/src/data/int/order/basic.lean b/src/data/int/order/basic.lean index 4f84a3537d7aa..2b426b0dd1aaa 100644 --- a/src/data/int/order/basic.lean +++ b/src/data/int/order/basic.lean @@ -525,7 +525,7 @@ theorem neg_div_of_dvd : ∀ {a b : ℤ} (H : b ∣ a), -a / b = -(a / b) lemma sub_div_of_dvd (a : ℤ) {b c : ℤ} (hcb : c ∣ b) : (a - b) / c = a / c - b / c := begin - rw [sub_eq_add_neg, sub_eq_add_neg, int.add_div_of_dvd_right ((dvd_neg c b).mpr hcb)], + rw [sub_eq_add_neg, sub_eq_add_neg, int.add_div_of_dvd_right hcb.neg_right], congr, exact neg_div_of_dvd hcb, end diff --git a/src/data/nat/gcd/basic.lean b/src/data/nat/gcd/basic.lean index c06cb1589bfbd..1755c5bc55313 100644 --- a/src/data/nat/gcd/basic.lean +++ b/src/data/nat/gcd/basic.lean @@ -266,6 +266,9 @@ dvd_antisymm theorem lcm_ne_zero {m n : ℕ} (hm : m ≠ 0) (hn : n ≠ 0) : lcm m n ≠ 0 := by { intro h, simpa [h, hm, hn] using gcd_mul_lcm m n, } +lemma lcm_pos {m n : ℕ} : 0 < m → 0 < n → 0 < m.lcm n := +by { simp_rw pos_iff_ne_zero, exact lcm_ne_zero } + /-! ### `coprime` diff --git a/src/data/nat/order/basic.lean b/src/data/nat/order/basic.lean index e1a9995b9225f..fcd6864736ece 100644 --- a/src/data/nat/order/basic.lean +++ b/src/data/nat/order/basic.lean @@ -65,7 +65,7 @@ instance : canonically_linear_ordered_add_monoid ℕ := { .. (infer_instance : canonically_ordered_add_monoid ℕ), .. nat.linear_order } -variables {m n k l : ℕ} +variables {a b m n k l : ℕ} namespace nat /-! ### Equalities and inequalities involving zero and one -/ @@ -94,7 +94,6 @@ lemma eq_zero_of_mul_le (hb : 2 ≤ n) (h : n * m ≤ m) : m = 0 := eq_zero_of_double_le $ le_trans (nat.mul_le_mul_right _ hb) h lemma zero_max : max 0 n = n := max_eq_right (zero_le _) - @[simp] lemma min_eq_zero_iff : min m n = 0 ↔ m = 0 ∨ n = 0 := begin split, @@ -243,17 +242,6 @@ lemma sub_succ' (m n : ℕ) : m - n.succ = m - n - 1 := rfl /-! ### `mul` -/ -lemma mul_eq_one_iff : ∀ {m n : ℕ}, m * n = 1 ↔ m = 1 ∧ n = 1 -| 0 0 := dec_trivial -| 0 1 := dec_trivial -| 1 0 := dec_trivial -| (m+2) 0 := by simp -| 0 (n+2) := by simp -| (m+1) (n+1) := ⟨ - λ h, by simp only [add_mul, mul_add, mul_add, one_mul, mul_one, - (add_assoc _ _ _).symm, nat.succ_inj', add_eq_zero_iff] at h; simp [h.1.2, h.2], - λ h, by simp only [h, mul_one]⟩ - lemma succ_mul_pos (m : ℕ) (hn : 0 < n) : 0 < (succ m) * n := mul_pos (succ_pos m) hn theorem mul_self_le_mul_self (h : m ≤ n) : m * m ≤ n * n := mul_le_mul h h (zero_le _) (zero_le _) diff --git a/src/data/nat/order/lemmas.lean b/src/data/nat/order/lemmas.lean index 95e76cfe69ee5..84cddcf68ccce 100644 --- a/src/data/nat/order/lemmas.lean +++ b/src/data/nat/order/lemmas.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn, Leonardo de Moura, Jeremy Avigad, Mario Carneiro -/ import data.nat.order.basic +import data.nat.units import data.set.basic import algebra.ring.divisibility import algebra.group_with_zero.divisibility diff --git a/src/group_theory/perm/cycle/basic.lean b/src/group_theory/perm/cycle/basic.lean index 24d92b6a295be..924ad2f62a9ac 100644 --- a/src/group_theory/perm/cycle/basic.lean +++ b/src/group_theory/perm/cycle/basic.lean @@ -736,7 +736,7 @@ lemma is_cycle_on.zpow_apply_eq {s : finset α} (hf : f.is_cycle_on s) (ha : a ∀ {n : ℤ}, (f ^ n) a = a ↔ (s.card : ℤ) ∣ n | (int.of_nat n) := (hf.pow_apply_eq ha).trans int.coe_nat_dvd.symm | (int.neg_succ_of_nat n) := by { rw [zpow_neg_succ_of_nat, ←inv_pow], - exact (hf.inv.pow_apply_eq ha).trans ((dvd_neg _ _).trans int.coe_nat_dvd).symm } + exact (hf.inv.pow_apply_eq ha).trans (dvd_neg.trans int.coe_nat_dvd).symm } lemma is_cycle_on.pow_apply_eq_pow_apply {s : finset α} (hf : f.is_cycle_on s) (ha : a ∈ s) {m n : ℕ} : (f ^ m) a = (f ^ n) a ↔ m ≡ n [MOD s.card] := diff --git a/src/number_theory/arithmetic_function.lean b/src/number_theory/arithmetic_function.lean index 332b45f5f2cd3..6c7ead1903114 100644 --- a/src/number_theory/arithmetic_function.lean +++ b/src/number_theory/arithmetic_function.lean @@ -656,7 +656,7 @@ begin rcases eq_or_ne m 1 with rfl | hm', { simp }, rw [one_apply_ne, one_apply_ne hm', zero_mul], - rw [ne.def, nat.mul_eq_one_iff, not_and_distrib], + rw [ne.def, mul_eq_one, not_and_distrib], exact or.inl hm' end⟩ diff --git a/src/number_theory/divisors.lean b/src/number_theory/divisors.lean index 011394390da6d..f069d24981d60 100644 --- a/src/number_theory/divisors.lean +++ b/src/number_theory/divisors.lean @@ -185,7 +185,7 @@ lemma divisors_antidiagonal_zero : divisors_antidiagonal 0 = ∅ := by { ext, si @[simp] lemma divisors_antidiagonal_one : divisors_antidiagonal 1 = {(1,1)} := -by { ext, simp [nat.mul_eq_one_iff, prod.ext_iff], } +by { ext, simp [mul_eq_one, prod.ext_iff], } @[simp] lemma swap_mem_divisors_antidiagonal {x : ℕ × ℕ} : x.swap ∈ divisors_antidiagonal n ↔ x ∈ divisors_antidiagonal n := diff --git a/src/number_theory/multiplicity.lean b/src/number_theory/multiplicity.lean index 8df0b65f91b6c..5823e572d478f 100644 --- a/src/number_theory/multiplicity.lean +++ b/src/number_theory/multiplicity.lean @@ -45,7 +45,7 @@ end lemma dvd_geom_sum₂_iff_of_dvd_sub' {x y p : R} (h : p ∣ x - y) : p ∣ ∑ i in range n, x ^ i * y ^ (n - 1 - i) ↔ p ∣ n * x ^ (n - 1) := -by rw [geom_sum₂_comm, dvd_geom_sum₂_iff_of_dvd_sub]; simpa using (dvd_neg _ _).mpr h +by rw [geom_sum₂_comm, dvd_geom_sum₂_iff_of_dvd_sub]; simpa using h.neg_right lemma dvd_geom_sum₂_self {x y : R} (h : ↑n ∣ x - y) : ↑n ∣ ∑ i in range n, x ^ i * y ^ (n - 1 - i):= (dvd_geom_sum₂_iff_of_dvd_sub h).mpr (dvd_mul_right _ _) diff --git a/src/number_theory/pythagorean_triples.lean b/src/number_theory/pythagorean_triples.lean index da59fbbe9b361..dfbe9fe522e63 100644 --- a/src/number_theory/pythagorean_triples.lean +++ b/src/number_theory/pythagorean_triples.lean @@ -129,8 +129,8 @@ begin { -- x even, y even exfalso, apply nat.not_coprime_of_dvd_of_dvd (dec_trivial : 1 < 2) _ _ hc, - { apply int.dvd_nat_abs_of_of_nat_dvd, apply int.dvd_of_mod_eq_zero hx }, - { apply int.dvd_nat_abs_of_of_nat_dvd, apply int.dvd_of_mod_eq_zero hy } }, + { apply int.coe_nat_dvd_left.1, apply int.dvd_of_mod_eq_zero hx }, + { apply int.coe_nat_dvd_left.1, apply int.dvd_of_mod_eq_zero hy } }, { left, exact ⟨hx, hy⟩ }, -- x even, y odd { right, exact ⟨hx, hy⟩ }, -- x odd, y even { -- x odd, y odd @@ -334,8 +334,7 @@ begin apply mt (int.dvd_gcd (int.coe_nat_dvd_left.mpr hpm)) hnp, apply (or_self _).mp, apply int.prime.dvd_mul' hp, rw (by ring : n * n = - (m ^ 2 - n ^ 2) + m * m), - apply dvd_add (dvd_neg_of_dvd hp1), - exact dvd_mul_of_dvd_left (int.coe_nat_dvd_left.mpr hpm) m }, + exact hp1.neg_right.add ((int.coe_nat_dvd_left.2 hpm).mul_right _) }, rw int.gcd_comm at hnp, apply mt (int.dvd_gcd (int.coe_nat_dvd_left.mpr hpn)) hnp, apply (or_self _).mp, apply int.prime.dvd_mul' hp, diff --git a/src/number_theory/zsqrtd/basic.lean b/src/number_theory/zsqrtd/basic.lean index ded1c18c7e42a..339fa34f9b9fe 100644 --- a/src/number_theory/zsqrtd/basic.lean +++ b/src/number_theory/zsqrtd/basic.lean @@ -388,7 +388,7 @@ lemma norm_eq_one_iff {x : ℤ√d} : x.norm.nat_abs = 1 ↔ is_unit x := λ h, let ⟨y, hy⟩ := is_unit_iff_dvd_one.1 h in begin have := congr_arg (int.nat_abs ∘ norm) hy, rw [function.comp_app, function.comp_app, norm_mul, int.nat_abs_mul, - norm_one, int.nat_abs_one, eq_comm, nat.mul_eq_one_iff] at this, + norm_one, int.nat_abs_one, eq_comm, mul_eq_one] at this, exact this.1 end⟩ diff --git a/src/number_theory/zsqrtd/gaussian_int.lean b/src/number_theory/zsqrtd/gaussian_int.lean index 363cfc9dad6ca..66a9a7c1451b3 100644 --- a/src/number_theory/zsqrtd/gaussian_int.lean +++ b/src/number_theory/zsqrtd/gaussian_int.lean @@ -243,7 +243,7 @@ hp.1.eq_two_or_odd.elim (λ hx0, (show (1 : ℤ) ≠ 0, from dec_trivial) $ by simpa [hx0] using congr_arg zsqrtd.im hx), have hpu : ¬ is_unit (p : ℤ[i]), from mt norm_eq_one_iff.2 - (by rw [norm_nat_cast, int.nat_abs_mul, nat.mul_eq_one_iff]; + (by rw [norm_nat_cast, int.nat_abs_mul, mul_eq_one]; exact λ h, (ne_of_lt hp.1.one_lt).symm h.1), obtain ⟨y, hy⟩ := hpk, have := hpi.2.2 ⟨k, 1⟩ ⟨k, -1⟩ ⟨y, by rw [← hkmul, ← nat.cast_mul p, ← hy]; simp⟩, @@ -253,7 +253,7 @@ hp.1.eq_two_or_odd.elim lemma sq_add_sq_of_nat_prime_of_not_irreducible (p : ℕ) [hp : fact p.prime] (hpi : ¬irreducible (p : ℤ[i])) : ∃ a b, a^2 + b^2 = p := have hpu : ¬ is_unit (p : ℤ[i]), from mt norm_eq_one_iff.2 $ - by rw [norm_nat_cast, int.nat_abs_mul, nat.mul_eq_one_iff]; + by rw [norm_nat_cast, int.nat_abs_mul, mul_eq_one]; exact λ h, (ne_of_lt hp.1.one_lt).symm h.1, have hab : ∃ a b, (p : ℤ[i]) = a * b ∧ ¬ is_unit a ∧ ¬ is_unit b, by simpa [irreducible_iff, hpu, not_forall, not_or_distrib] using hpi, diff --git a/src/ring_theory/multiplicity.lean b/src/ring_theory/multiplicity.lean index f51eca9e5324d..a0bfb28b60031 100644 --- a/src/ring_theory/multiplicity.lean +++ b/src/ring_theory/multiplicity.lean @@ -330,8 +330,8 @@ variables [ring α] [decidable_rel ((∣) : α → α → Prop)] @[simp] protected lemma neg (a b : α) : multiplicity a (-b) = multiplicity a b := part.ext' (by simp only [multiplicity, part_enat.find, dvd_neg]) (λ h₁ h₂, part_enat.coe_inj.1 (by rw [part_enat.coe_get]; exact - eq.symm (unique ((dvd_neg _ _).2 (pow_multiplicity_dvd _)) - (mt (dvd_neg _ _).1 (is_greatest' _ (lt_succ_self _)))))) + eq.symm (unique (pow_multiplicity_dvd _).neg_right + (mt dvd_neg.1 (is_greatest' _ (lt_succ_self _)))))) theorem int.nat_abs (a : ℕ) (b : ℤ) : multiplicity a b.nat_abs = multiplicity (a : ℤ) b := begin @@ -346,8 +346,8 @@ begin apply le_antisymm, { apply part_enat.le_of_lt_add_one, cases part_enat.ne_top_iff.mp (part_enat.ne_top_of_lt h) with k hk, - rw [hk], rw_mod_cast [multiplicity_lt_iff_neg_dvd], intro h_dvd, - rw [← dvd_add_iff_right] at h_dvd, + rw [hk], rw_mod_cast [multiplicity_lt_iff_neg_dvd, dvd_add_right], + intro h_dvd, apply multiplicity.is_greatest _ h_dvd, rw [hk], apply_mod_cast nat.lt_succ_self, rw [pow_dvd_iff_le_multiplicity, nat.cast_add, ← hk, nat.cast_one], exact part_enat.add_one_le_of_lt h }, diff --git a/src/tactic/omega/eq_elim.lean b/src/tactic/omega/eq_elim.lean index d73846762eb16..20bf778edbcfd 100644 --- a/src/tactic/omega/eq_elim.lean +++ b/src/tactic/omega/eq_elim.lean @@ -125,7 +125,7 @@ begin simp only [term.val, mul_add, add_mul, m, a_n], ring }, cases h4 with c h4, - rw [dvd_add_iff_right (dvd_mul_right m c), h4, ← h1], + rw [←dvd_add_right (dvd_mul_right m c), h4, ← h1], apply dvd_zero }, apply calc v n = -(m * sgm v b as n) + (symmod b m) + @@ -375,7 +375,7 @@ lemma sat_eq_elim : have h3 : 0 = b + coeffs.val v as := h1.left _ (or.inl rfl), have h4 : i ∣ coeffs.val v as := coeffs.dvd_val h2.right, have h5 : i ∣ b + coeffs.val v as := by { rw ← h3, apply dvd_zero }, - rw ← dvd_add_iff_left h4 at h5, apply h2.left h5 }, + rw dvd_add_left h4 at h5, apply h2.left h5 }, rw if_neg h2, apply sat_empty end | (ee.factor i::es) ((b,as)::eqs, les) h1 := From b2a5f0d6fc79f4aa24586177a8d33b20daf3aea5 Mon Sep 17 00:00:00 2001 From: Moritz Doll Date: Wed, 5 Apr 2023 10:55:53 +0000 Subject: [PATCH 0777/1291] feat(analysis/schwartz_space): construction of continuous linear maps (#18651) This PR has two parts: (a) create two definitions to construct continuous linear maps (the point being that one wants to reuse the boundedness estimate for both the well-posedness and the continuity) and (b) drastically reduce the boilerplate for `schwartz_map.fderiv_clm`. --- src/analysis/schwartz_space.lean | 118 ++++++++++++++++++------------- 1 file changed, 68 insertions(+), 50 deletions(-) diff --git a/src/analysis/schwartz_space.lean b/src/analysis/schwartz_space.lean index 0a2d35e132e9a..61f5f4ace9708 100644 --- a/src/analysis/schwartz_space.lean +++ b/src/analysis/schwartz_space.lean @@ -54,7 +54,7 @@ Schwartz space, tempered distributions noncomputable theory -variables {𝕜 𝕜' E F : Type*} +variables {𝕜 𝕜' D E F G : Type*} variables [normed_add_comm_group E] [normed_space ℝ E] variables [normed_add_comm_group F] [normed_space ℝ F] @@ -458,63 +458,81 @@ instance : topological_space.first_countable_topology (𝓢(E, F)) := end topology -section fderiv +section clm + +/-! ### Construction of continuous linear maps between Schwartz spaces -/ + +variables [normed_field 𝕜] [normed_field 𝕜'] +variables [normed_add_comm_group D] [normed_space ℝ D] +variables [normed_space 𝕜 E] [smul_comm_class ℝ 𝕜 E] +variables [normed_add_comm_group G] [normed_space ℝ G] [normed_space 𝕜' G] [smul_comm_class ℝ 𝕜' G] +variables {σ : 𝕜 →+* 𝕜'} + +/-- Create a semilinear map between Schwartz spaces. + +Note: This is a helper definition for `mk_clm`. -/ +def mk_lm (A : (D → E) → (F → G)) + (hadd : ∀ (f g : 𝓢(D, E)) x, A (f + g) x = A f x + A g x) + (hsmul : ∀ (a : 𝕜) (f : 𝓢(D, E)) x, A (a • f) x = σ a • A f x) + (hsmooth : ∀ (f : 𝓢(D, E)), cont_diff ℝ ⊤ (A f)) + (hbound : ∀ (n : ℕ × ℕ), ∃ (s : finset (ℕ × ℕ)) (C : ℝ) (hC : 0 ≤ C), ∀ (f : 𝓢(D, E)) (x : F), + ‖x‖ ^ n.fst * ‖iterated_fderiv ℝ n.snd (A f) x‖ ≤ C * s.sup (schwartz_seminorm_family 𝕜 D E) f) : + 𝓢(D, E) →ₛₗ[σ] 𝓢(F, G) := +{ to_fun := λ f, + { to_fun := A f, + smooth' := hsmooth f, + decay' := + begin + intros k n, + rcases hbound ⟨k, n⟩ with ⟨s, C, hC, h⟩, + exact ⟨C * (s.sup (schwartz_seminorm_family 𝕜 D E)) f, h f⟩, + end, }, + map_add' := λ f g, ext (hadd f g), + map_smul' := λ a f, ext (hsmul a f), } + +/-- Create a continuous semilinear map between Schwartz spaces. + +For an example of using this definition, see `fderiv_clm`. -/ +def mk_clm [ring_hom_isometric σ] (A : (D → E) → (F → G)) + (hadd : ∀ (f g : 𝓢(D, E)) x, A (f + g) x = A f x + A g x) + (hsmul : ∀ (a : 𝕜) (f : 𝓢(D, E)) x, A (a • f) x = σ a • A f x) + (hsmooth : ∀ (f : 𝓢(D, E)), cont_diff ℝ ⊤ (A f)) + (hbound : ∀ (n : ℕ × ℕ), ∃ (s : finset (ℕ × ℕ)) (C : ℝ) (hC : 0 ≤ C), ∀ (f : 𝓢(D, E)) (x : F), + ‖x‖ ^ n.fst * ‖iterated_fderiv ℝ n.snd (A f) x‖ ≤ C * s.sup (schwartz_seminorm_family 𝕜 D E) f) : + 𝓢(D, E) →SL[σ] 𝓢(F, G) := +{ cont := + begin + change continuous (mk_lm A hadd hsmul hsmooth hbound : 𝓢(D, E) →ₛₗ[σ] 𝓢(F, G)), + refine seminorm.continuous_from_bounded (schwartz_with_seminorms 𝕜 D E) + (schwartz_with_seminorms 𝕜' F G) _ (λ n, _), + rcases hbound n with ⟨s, C, hC, h⟩, + refine ⟨s, ⟨C, hC⟩, (λ f, _)⟩, + simp only [seminorm.comp_apply, seminorm.smul_apply, nnreal.smul_def, algebra.id.smul_eq_mul, + subtype.coe_mk], + exact (mk_lm A hadd hsmul hsmooth hbound f).seminorm_le_bound 𝕜' n.1 n.2 (by positivity) (h f), + end, + to_linear_map := mk_lm A hadd hsmul hsmooth hbound } -/-! ### Derivatives of Schwartz functions -/ +end clm -variables {E F} - -/-- The derivative of a Schwartz function as a Schwartz function with values in the -continuous linear maps `E→L[ℝ] F`. -/ -@[protected] def fderiv (f : 𝓢(E, F)) : 𝓢(E, E →L[ℝ] F) := -{ to_fun := fderiv ℝ f, - smooth' := (cont_diff_top_iff_fderiv.mp f.smooth').2, - decay' := - begin - intros k n, - cases f.decay' k (n+1) with C hC, - use C, - intros x, - rw norm_iterated_fderiv_fderiv, - exact hC x, - end } +section fderiv -@[simp, norm_cast] lemma coe_fderiv (f : 𝓢(E, F)) : ⇑f.fderiv = fderiv ℝ f := rfl -@[simp] lemma fderiv_apply (f : 𝓢(E, F)) (x : E) : f.fderiv x = fderiv ℝ f x := rfl +/-! ### Derivatives of Schwartz functions -/ variables (𝕜) variables [is_R_or_C 𝕜] [normed_space 𝕜 F] [smul_comm_class ℝ 𝕜 F] -/-- The derivative on Schwartz space as a linear map. -/ -def fderiv_lm : 𝓢(E, F) →ₗ[𝕜] 𝓢(E, E →L[ℝ] F) := -{ to_fun := schwartz_map.fderiv, - map_add' := λ f g, ext $ λ _, fderiv_add - f.differentiable.differentiable_at - g.differentiable.differentiable_at, - map_smul' := λ a f, ext $ λ _, fderiv_const_smul f.differentiable.differentiable_at a } - -@[simp, norm_cast] lemma fderiv_lm_apply (f : 𝓢(E, F)) : fderiv_lm 𝕜 f = schwartz_map.fderiv f := -rfl - -/-- The derivative on Schwartz space as a continuous linear map. -/ +/-- The real derivative on Schwartz space as a continuous `𝕜`-linear map. -/ def fderiv_clm : 𝓢(E, F) →L[𝕜] 𝓢(E, E →L[ℝ] F) := -{ cont := - begin - change continuous (fderiv_lm 𝕜 : 𝓢(E, F) →ₗ[𝕜] 𝓢(E, E →L[ℝ] F)), - refine seminorm.continuous_from_bounded (schwartz_with_seminorms 𝕜 E F) - (schwartz_with_seminorms 𝕜 E (E →L[ℝ] F)) _ _, - rintros ⟨k, n⟩, - use [{⟨k, n+1⟩}, 1], - intros f, - simp only [schwartz_seminorm_family_apply, seminorm.comp_apply, finset.sup_singleton, one_smul], - refine (fderiv_lm 𝕜 f).seminorm_le_bound 𝕜 k n (by positivity) _, - intros x, - rw [fderiv_lm_apply, coe_fderiv, norm_iterated_fderiv_fderiv], - exact f.le_seminorm 𝕜 k (n+1) x, - end, - to_linear_map := fderiv_lm 𝕜 } - -@[simp, norm_cast] lemma fderiv_clm_apply (f : 𝓢(E, F)) : fderiv_clm 𝕜 f = schwartz_map.fderiv f := +mk_clm (fderiv ℝ) + (λ f g _, fderiv_add f.differentiable.differentiable_at g.differentiable.differentiable_at) + (λ a f _, fderiv_const_smul f.differentiable.differentiable_at a) + (λ f, (cont_diff_top_iff_fderiv.mp f.smooth').2) + (λ ⟨k, n⟩, ⟨{⟨k, n+1⟩}, 1, zero_le_one, λ f x, by simpa only [schwartz_seminorm_family_apply, + seminorm.comp_apply, finset.sup_singleton, one_smul, norm_iterated_fderiv_fderiv, one_mul] + using f.le_seminorm 𝕜 k (n+1) x⟩) + +@[simp] lemma fderiv_clm_apply (f : 𝓢(E, F)) (x : E) : fderiv_clm 𝕜 f x = fderiv ℝ f x := rfl end fderiv From 0ac3057eb6231d2c8dfcd46767cf4a166961c0f1 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Wed, 5 Apr 2023 10:55:54 +0000 Subject: [PATCH 0778/1291] feat(field_theory/is_alg_closed/algebraic_closure): add instances normal and is_galois for algebraic closure (#18730) --- src/field_theory/galois.lean | 10 +++++++++- .../is_alg_closed/algebraic_closure.lean | 2 +- src/field_theory/is_alg_closed/basic.lean | 16 ++++++++++++++-- 3 files changed, 24 insertions(+), 4 deletions(-) diff --git a/src/field_theory/galois.lean b/src/field_theory/galois.lean index 18b8cf16992b7..c760d259fb005 100644 --- a/src/field_theory/galois.lean +++ b/src/field_theory/galois.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Thomas Browning, Patrick Lutz -/ -import field_theory.normal +import field_theory.is_alg_closed.algebraic_closure import field_theory.primitive_element import field_theory.fixed import group_theory.group_action.fixing_subgroup @@ -424,3 +424,11 @@ end end is_galois end galois_equivalent_definitions + +section is_alg_closure + +@[priority 100] +instance is_alg_closure.is_galois (k K : Type*) [field k] [field K] [algebra k K] + [is_alg_closure k K] [char_zero k] : is_galois k K := { } + +end is_alg_closure diff --git a/src/field_theory/is_alg_closed/algebraic_closure.lean b/src/field_theory/is_alg_closed/algebraic_closure.lean index 80769023bd040..6435bb6ed3c2a 100644 --- a/src/field_theory/is_alg_closed/algebraic_closure.lean +++ b/src/field_theory/is_alg_closed/algebraic_closure.lean @@ -5,7 +5,7 @@ Authors: Kenny Lau -/ import algebra.direct_limit import field_theory.is_alg_closed.basic -import field_theory.splitting_field + /-! # Algebraic Closure diff --git a/src/field_theory/is_alg_closed/basic.lean b/src/field_theory/is_alg_closed/basic.lean index b23f594b1fee5..d04ad10dc000c 100644 --- a/src/field_theory/is_alg_closed/basic.lean +++ b/src/field_theory/is_alg_closed/basic.lean @@ -4,9 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau -/ +import field_theory.normal import field_theory.perfect_closure -import field_theory.separable -import ring_theory.adjoin.field import ring_theory.localization.integral /-! @@ -170,6 +169,19 @@ theorem is_alg_closure_iff (K : Type v) [field K] [algebra k K] : is_alg_closure k K ↔ is_alg_closed K ∧ algebra.is_algebraic k K := ⟨λ h, ⟨h.1, h.2⟩, λ h, ⟨h.1, h.2⟩⟩ +@[priority 100] +instance is_alg_closure.normal (R K : Type*) [field R] [field K] [algebra R K] [is_alg_closure R K]: + normal R K := +⟨is_alg_closure.algebraic, λ _, + @is_alg_closed.splits_codomain _ _ _ (is_alg_closure.alg_closed R) _ _ _⟩ + +@[priority 100] +instance is_alg_closure.separable (R K : Type*) [field R] [field K] [algebra R K] +[is_alg_closure R K] [char_zero R] : + is_separable R K := +⟨λ _, is_algebraic_iff_is_integral.mp (is_alg_closure.algebraic _), λ _, (minpoly.irreducible + (is_algebraic_iff_is_integral.mp (is_alg_closure.algebraic _))).separable⟩ + namespace lift /- In this section, the homomorphism from any algebraic extension into an algebraically From 039a089d2a4b93c761b234f3e5f5aeb752bac60f Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 5 Apr 2023 10:55:55 +0000 Subject: [PATCH 0779/1291] refactor(linear_algebra/dimension): use `rank` in lemma names instead of `dim` (#18741) The `dim` name is left from the previous name of the function, `vector_space.dim`. When that was merged with `module.rank` in #7322, we left renaming the existing lemmas as future work. This commit was made by replacing `(\b|(?<=_))dim(\b|(?=_))` with `rank` in the `dimension` and `finite_dimensional` files, and then manually fixing downstream breakages; it's important not to rename `power_basis.dim` at the same time! Deciding whether to move some of these to the `module` namespace is left as future work, the diff is already big enough. --- archive/sensitivity.lean | 18 +- src/algebra/linear_recurrence.lean | 4 +- src/algebra/quaternion.lean | 10 +- .../inner_product_space/projection.lean | 2 +- src/data/complex/module.lean | 12 +- src/field_theory/adjoin.lean | 32 +- src/field_theory/finite/polynomial.lean | 13 +- src/field_theory/finiteness.lean | 16 +- src/field_theory/fixed.lean | 14 +- src/field_theory/intermediate_field.lean | 2 +- src/field_theory/krull_topology.lean | 2 +- src/field_theory/mv_polynomial.lean | 4 +- src/field_theory/tower.lean | 12 +- .../affine_space/finite_dimensional.lean | 24 +- src/linear_algebra/bilinear_form.lean | 2 +- src/linear_algebra/dimension.lean | 336 +++++++++--------- src/linear_algebra/dual.lean | 8 +- src/linear_algebra/finite_dimensional.lean | 162 ++++----- src/linear_algebra/finrank.lean | 60 ++-- src/linear_algebra/free_algebra.lean | 4 +- .../free_module/finite/matrix.lean | 2 +- .../free_module/finite/rank.lean | 2 +- src/linear_algebra/free_module/rank.lean | 16 +- src/linear_algebra/matrix/diagonal.lean | 4 +- src/linear_algebra/matrix/to_lin.lean | 2 +- src/number_theory/ramification_inertia.lean | 24 +- 26 files changed, 396 insertions(+), 391 deletions(-) diff --git a/archive/sensitivity.lean b/archive/sensitivity.lean index 5b96d7ddbffa0..6b3b3e9f6804f 100644 --- a/archive/sensitivity.lean +++ b/archive/sensitivity.lean @@ -230,7 +230,7 @@ since this cardinal is finite, as a natural number in `finrank_V` -/ lemma dim_V : module.rank ℝ (V n) = 2^n := have module.rank ℝ (V n) = (2^n : ℕ), - by { rw [dim_eq_card_basis (dual_bases_e_ε _).basis, Q.card]; apply_instance }, + by { rw [rank_eq_card_basis (dual_bases_e_ε _).basis, Q.card]; apply_instance }, by assumption_mod_cast instance : finite_dimensional ℝ (V n) := @@ -238,7 +238,7 @@ finite_dimensional.of_fintype_basis (dual_bases_e_ε _).basis lemma finrank_V : finrank ℝ (V n) = 2^n := have _ := @dim_V n, -by rw ←finrank_eq_dim at this; assumption_mod_cast +by rw ←finrank_eq_rank at this; assumption_mod_cast /-! ### The linear map -/ @@ -359,25 +359,25 @@ begin let img := (g m).range, suffices : 0 < dim (W ⊓ img), { simp only [exists_prop], - exact_mod_cast exists_mem_ne_zero_of_dim_pos this }, + exact_mod_cast exists_mem_ne_zero_of_rank_pos this }, have dim_le : dim (W ⊔ img) ≤ 2^(m + 1), - { convert ← dim_submodule_le (W ⊔ img), + { convert ← rank_submodule_le (W ⊔ img), apply dim_V }, have dim_add : dim (W ⊔ img) + dim (W ⊓ img) = dim W + 2^m, - { convert ← dim_sup_add_dim_inf_eq W img, - rw ← dim_eq_of_injective (g m) g_injective, + { convert ← rank_sup_add_rank_inf_eq W img, + rw ← rank_eq_of_injective (g m) g_injective, apply dim_V }, have dimW : dim W = card H, { have li : linear_independent ℝ (H.restrict e), { convert (dual_bases_e_ε _).basis.linear_independent.comp _ subtype.val_injective, rw (dual_bases_e_ε _).coe_basis }, - have hdW := dim_span li, + have hdW := rank_span li, rw set.range_restrict at hdW, convert hdW, rw [← (dual_bases_e_ε _).coe_basis, cardinal.mk_image_eq (dual_bases_e_ε _).basis.injective, cardinal.mk_fintype] }, - rw ← finrank_eq_dim ℝ at ⊢ dim_le dim_add dimW, - rw [← finrank_eq_dim ℝ, ← finrank_eq_dim ℝ] at dim_add, + rw ← finrank_eq_rank ℝ at ⊢ dim_le dim_add dimW, + rw [← finrank_eq_rank ℝ, ← finrank_eq_rank ℝ] at dim_add, norm_cast at ⊢ dim_le dim_add dimW, rw pow_succ' at dim_le, rw set.to_finset_card at hH, diff --git a/src/algebra/linear_recurrence.lean b/src/algebra/linear_recurrence.lean index 14461e059c4cd..e1e570166c3ae 100644 --- a/src/algebra/linear_recurrence.lean +++ b/src/algebra/linear_recurrence.lean @@ -174,10 +174,10 @@ section strong_rank_condition variables {α : Type*} [comm_ring α] [strong_rank_condition α] (E : linear_recurrence α) /-- The dimension of `E.sol_space` is `E.order`. -/ -lemma sol_space_dim : module.rank α E.sol_space = E.order := +lemma sol_space_rank : module.rank α E.sol_space = E.order := begin letI := nontrivial_of_invariant_basis_number α, - exact @dim_fin_fun α _ _ E.order ▸ E.to_init.dim_eq + exact @rank_fin_fun α _ _ E.order ▸ E.to_init.rank_eq end end strong_rank_condition diff --git a/src/algebra/quaternion.lean b/src/algebra/quaternion.lean index c5606fcbeb912..df53f8c4b7f83 100644 --- a/src/algebra/quaternion.lean +++ b/src/algebra/quaternion.lean @@ -309,13 +309,13 @@ basis.of_equiv_fun $ linear_equiv_tuple c₁ c₂ instance : module.finite R ℍ[R, c₁, c₂] := module.finite.of_basis (basis_one_i_j_k c₁ c₂) instance : module.free R ℍ[R, c₁, c₂] := module.free.of_basis (basis_one_i_j_k c₁ c₂) -lemma dim_eq_four [strong_rank_condition R] : module.rank R ℍ[R, c₁, c₂] = 4 := -by { rw [dim_eq_card_basis (basis_one_i_j_k c₁ c₂), fintype.card_fin], norm_num } +lemma rank_eq_four [strong_rank_condition R] : module.rank R ℍ[R, c₁, c₂] = 4 := +by { rw [rank_eq_card_basis (basis_one_i_j_k c₁ c₂), fintype.card_fin], norm_num } lemma finrank_eq_four [strong_rank_condition R] : finite_dimensional.finrank R ℍ[R, c₁, c₂] = 4 := have cardinal.to_nat 4 = 4, by rw [←cardinal.to_nat_cast 4, nat.cast_bit0, nat.cast_bit0, nat.cast_one], -by rw [finite_dimensional.finrank, dim_eq_four, this] +by rw [finite_dimensional.finrank, rank_eq_four, this] end @@ -631,8 +631,8 @@ lemma smul_coe : x • (y : ℍ[R]) = ↑(x * y) := quaternion_algebra.smul_coe instance : module.finite R ℍ[R] := quaternion_algebra.module.finite _ _ instance : module.free R ℍ[R] := quaternion_algebra.module.free _ _ -lemma dim_eq_four [strong_rank_condition R] : module.rank R ℍ[R] = 4 := -quaternion_algebra.dim_eq_four _ _ +lemma rank_eq_four [strong_rank_condition R] : module.rank R ℍ[R] = 4 := +quaternion_algebra.rank_eq_four _ _ lemma finrank_eq_four [strong_rank_condition R] : finite_dimensional.finrank R ℍ[R] = 4 := quaternion_algebra.finrank_eq_four _ _ diff --git a/src/analysis/inner_product_space/projection.lean b/src/analysis/inner_product_space/projection.lean index cc4c578cbd439..c23076664697c 100644 --- a/src/analysis/inner_product_space/projection.lean +++ b/src/analysis/inner_product_space/projection.lean @@ -1047,7 +1047,7 @@ lemma submodule.finrank_add_inf_finrank_orthogonal {K₁ K₂ : submodule 𝕜 E begin haveI := submodule.finite_dimensional_of_le h, haveI := proper_is_R_or_C 𝕜 K₁, - have hd := submodule.dim_sup_add_dim_inf_eq K₁ (K₁ᗮ ⊓ K₂), + have hd := submodule.rank_sup_add_rank_inf_eq K₁ (K₁ᗮ ⊓ K₂), rw [←inf_assoc, (submodule.orthogonal_disjoint K₁).eq_bot, bot_inf_eq, finrank_bot, submodule.sup_orthogonal_inf_of_complete_space h] at hd, rw add_zero at hd, diff --git a/src/data/complex/module.lean b/src/data/complex/module.lean index 332f57af117db..d3a67b8891069 100644 --- a/src/data/complex/module.lean +++ b/src/data/complex/module.lean @@ -157,11 +157,11 @@ instance : finite_dimensional ℝ ℂ := of_fintype_basis basis_one_I @[simp] lemma finrank_real_complex : finite_dimensional.finrank ℝ ℂ = 2 := by rw [finrank_eq_card_basis basis_one_I, fintype.card_fin] -@[simp] lemma dim_real_complex : module.rank ℝ ℂ = 2 := -by simp [← finrank_eq_dim, finrank_real_complex] +@[simp] lemma rank_real_complex : module.rank ℝ ℂ = 2 := +by simp [← finrank_eq_rank, finrank_real_complex] -lemma {u} dim_real_complex' : cardinal.lift.{u} (module.rank ℝ ℂ) = 2 := -by simp [← finrank_eq_dim, finrank_real_complex, bit0] +lemma {u} rank_real_complex' : cardinal.lift.{u} (module.rank ℝ ℂ) = 2 := +by simp [← finrank_eq_rank, finrank_real_complex, bit0] /-- `fact` version of the dimension of `ℂ` over `ℝ`, locally useful in the definition of the circle. -/ @@ -199,10 +199,10 @@ instance finite_dimensional.complex_to_real (E : Type*) [add_comm_group E] [modu [finite_dimensional ℂ E] : finite_dimensional ℝ E := finite_dimensional.trans ℝ ℂ E -lemma dim_real_of_complex (E : Type*) [add_comm_group E] [module ℂ E] : +lemma rank_real_of_complex (E : Type*) [add_comm_group E] [module ℂ E] : module.rank ℝ E = 2 * module.rank ℂ E := cardinal.lift_inj.1 $ - by { rw [← dim_mul_dim' ℝ ℂ E, complex.dim_real_complex], simp [bit0] } + by { rw [← rank_mul_rank' ℝ ℂ E, complex.rank_real_complex], simp [bit0] } lemma finrank_real_of_complex (E : Type*) [add_comm_group E] [module ℂ E] : finite_dimensional.finrank ℝ E = 2 * finite_dimensional.finrank ℂ E := diff --git a/src/field_theory/adjoin.lean b/src/field_theory/adjoin.lean index 283e523f6afdb..91338e32a54dc 100644 --- a/src/field_theory/adjoin.lean +++ b/src/field_theory/adjoin.lean @@ -19,7 +19,7 @@ For example, `algebra.adjoin K {x}` might not include `x⁻¹`. ## Main results - `adjoin_adjoin_left`: adjoining S and then T is the same as adjoining `S ∪ T`. -- `bot_eq_top_of_dim_adjoin_eq_one`: if `F⟮x⟯` has dimension `1` over `F` for every `x` +- `bot_eq_top_of_rank_adjoin_eq_one`: if `F⟮x⟯` has dimension `1` over `F` for every `x` in `E` then `F = E` ## Notation @@ -528,30 +528,30 @@ adjoin_simple_eq_bot_iff.mpr (coe_int_mem ⊥ n) @[simp] lemma adjoin_nat (n : ℕ) : F⟮(n : E)⟯ = ⊥ := adjoin_simple_eq_bot_iff.mpr (coe_nat_mem ⊥ n) -section adjoin_dim +section adjoin_rank open finite_dimensional module variables {K L : intermediate_field F E} -@[simp] lemma dim_eq_one_iff : module.rank F K = 1 ↔ K = ⊥ := -by rw [← to_subalgebra_eq_iff, ← dim_eq_dim_subalgebra, - subalgebra.dim_eq_one_iff, bot_to_subalgebra] +@[simp] lemma rank_eq_one_iff : module.rank F K = 1 ↔ K = ⊥ := +by rw [← to_subalgebra_eq_iff, ← rank_eq_rank_subalgebra, + subalgebra.rank_eq_one_iff, bot_to_subalgebra] @[simp] lemma finrank_eq_one_iff : finrank F K = 1 ↔ K = ⊥ := by rw [← to_subalgebra_eq_iff, ← finrank_eq_finrank_subalgebra, subalgebra.finrank_eq_one_iff, bot_to_subalgebra] -@[simp] lemma dim_bot : module.rank F (⊥ : intermediate_field F E) = 1 := -by rw dim_eq_one_iff +@[simp] lemma rank_bot : module.rank F (⊥ : intermediate_field F E) = 1 := +by rw rank_eq_one_iff @[simp] lemma finrank_bot : finrank F (⊥ : intermediate_field F E) = 1 := by rw finrank_eq_one_iff -lemma dim_adjoin_eq_one_iff : module.rank F (adjoin F S) = 1 ↔ S ⊆ (⊥ : intermediate_field F E) := -iff.trans dim_eq_one_iff adjoin_eq_bot_iff +lemma rank_adjoin_eq_one_iff : module.rank F (adjoin F S) = 1 ↔ S ⊆ (⊥ : intermediate_field F E) := +iff.trans rank_eq_one_iff adjoin_eq_bot_iff -lemma dim_adjoin_simple_eq_one_iff : module.rank F F⟮α⟯ = 1 ↔ α ∈ (⊥ : intermediate_field F E) := -by { rw dim_adjoin_eq_one_iff, exact set.singleton_subset_iff } +lemma rank_adjoin_simple_eq_one_iff : module.rank F F⟮α⟯ = 1 ↔ α ∈ (⊥ : intermediate_field F E) := +by { rw rank_adjoin_eq_one_iff, exact set.singleton_subset_iff } lemma finrank_adjoin_eq_one_iff : finrank F (adjoin F S) = 1 ↔ S ⊆ (⊥ : intermediate_field F E) := iff.trans finrank_eq_one_iff adjoin_eq_bot_iff @@ -560,12 +560,12 @@ lemma finrank_adjoin_simple_eq_one_iff : finrank F F⟮α⟯ = 1 ↔ α ∈ (⊥ by { rw [finrank_adjoin_eq_one_iff], exact set.singleton_subset_iff } /-- If `F⟮x⟯` has dimension `1` over `F` for every `x ∈ E` then `F = E`. -/ -lemma bot_eq_top_of_dim_adjoin_eq_one (h : ∀ x : E, module.rank F F⟮x⟯ = 1) : +lemma bot_eq_top_of_rank_adjoin_eq_one (h : ∀ x : E, module.rank F F⟮x⟯ = 1) : (⊥ : intermediate_field F E) = ⊤ := begin ext, rw iff_true_right intermediate_field.mem_top, - exact dim_adjoin_simple_eq_one_iff.mp (h x), + exact rank_adjoin_simple_eq_one_iff.mp (h x), end lemma bot_eq_top_of_finrank_adjoin_eq_one (h : ∀ x : E, finrank F F⟮x⟯ = 1) : @@ -576,9 +576,9 @@ begin exact finrank_adjoin_simple_eq_one_iff.mp (h x), end -lemma subsingleton_of_dim_adjoin_eq_one (h : ∀ x : E, module.rank F F⟮x⟯ = 1) : +lemma subsingleton_of_rank_adjoin_eq_one (h : ∀ x : E, module.rank F F⟮x⟯ = 1) : subsingleton (intermediate_field F E) := -subsingleton_of_bot_eq_top (bot_eq_top_of_dim_adjoin_eq_one h) +subsingleton_of_bot_eq_top (bot_eq_top_of_rank_adjoin_eq_one h) lemma subsingleton_of_finrank_adjoin_eq_one (h : ∀ x : E, finrank F F⟮x⟯ = 1) : subsingleton (intermediate_field F E) := @@ -596,7 +596,7 @@ lemma subsingleton_of_finrank_adjoin_le_one [finite_dimensional F E] (h : ∀ x : E, finrank F F⟮x⟯ ≤ 1) : subsingleton (intermediate_field F E) := subsingleton_of_bot_eq_top (bot_eq_top_of_finrank_adjoin_le_one h) -end adjoin_dim +end adjoin_rank end adjoin_intermediate_field_lattice section adjoin_integral_element diff --git a/src/field_theory/finite/polynomial.lean b/src/field_theory/finite/polynomial.lean index 31bc5929653de..b3b4000b72398 100644 --- a/src/field_theory/finite/polynomial.lean +++ b/src/field_theory/finite/polynomial.lean @@ -177,13 +177,13 @@ end comm_ring variables [field K] -lemma dim_R [fintype σ] : module.rank K (R σ K) = fintype.card (σ → K) := +lemma rank_R [fintype σ] : module.rank K (R σ K) = fintype.card (σ → K) := calc module.rank K (R σ K) = module.rank K (↥{s : σ →₀ ℕ | ∀ (n : σ), s n ≤ fintype.card K - 1} →₀ K) : - linear_equiv.dim_eq + linear_equiv.rank_eq (finsupp.supported_equiv_finsupp {s : σ →₀ ℕ | ∀n:σ, s n ≤ fintype.card K - 1 }) ... = #{s : σ →₀ ℕ | ∀ (n : σ), s n ≤ fintype.card K - 1} : - by rw [finsupp.dim_eq, dim_self, mul_one] + by rw [finsupp.rank_eq, rank_self, mul_one] ... = #{s : σ → ℕ | ∀ (n : σ), s n < fintype.card K } : begin refine quotient.sound ⟨equiv.subtype_equiv finsupp.equiv_fun_on_finite $ assume f, _⟩, @@ -199,11 +199,12 @@ calc module.rank K (R σ K) = ... = fintype.card (σ → K) : cardinal.mk_fintype _ instance [finite σ] : finite_dimensional K (R σ K) := -by { casesI nonempty_fintype σ, exact is_noetherian.iff_fg.1 (is_noetherian.iff_dim_lt_aleph_0.mpr $ - by simpa only [dim_R] using cardinal.nat_lt_aleph_0 (fintype.card (σ → K))) } +by { casesI nonempty_fintype σ, + exact is_noetherian.iff_fg.1 (is_noetherian.iff_rank_lt_aleph_0.mpr $ + by simpa only [rank_R] using cardinal.nat_lt_aleph_0 (fintype.card (σ → K))) } lemma finrank_R [fintype σ] : finite_dimensional.finrank K (R σ K) = fintype.card (σ → K) := -finite_dimensional.finrank_eq_of_dim_eq (dim_R σ K) +finite_dimensional.finrank_eq_of_rank_eq (rank_R σ K) lemma range_evalᵢ [finite σ] : (evalᵢ σ K).range = ⊤ := begin diff --git a/src/field_theory/finiteness.lean b/src/field_theory/finiteness.lean index c67d38565a3c3..39f8634bdbc00 100644 --- a/src/field_theory/finiteness.lean +++ b/src/field_theory/finiteness.lean @@ -24,10 +24,10 @@ variables {K : Type u} {V : Type v} [division_ring K] [add_comm_group V] [module A module over a division ring is noetherian if and only if its dimension (as a cardinal) is strictly less than the first infinite cardinal `ℵ₀`. -/ -lemma iff_dim_lt_aleph_0 : is_noetherian K V ↔ module.rank K V < ℵ₀ := +lemma iff_rank_lt_aleph_0 : is_noetherian K V ↔ module.rank K V < ℵ₀ := begin let b := basis.of_vector_space K V, - rw [← b.mk_eq_dim'', lt_aleph_0_iff_set_finite], + rw [← b.mk_eq_rank'', lt_aleph_0_iff_set_finite], split, { introI, exact finite_of_linear_independent (basis.of_vector_space_index.linear_independent K V) }, @@ -42,15 +42,15 @@ variables (K V) /-- The dimension of a noetherian module over a division ring, as a cardinal, is strictly less than the first infinite cardinal `ℵ₀`. -/ -lemma dim_lt_aleph_0 : ∀ [is_noetherian K V], module.rank K V < ℵ₀ := -is_noetherian.iff_dim_lt_aleph_0.1 +lemma rank_lt_aleph_0 : ∀ [is_noetherian K V], module.rank K V < ℵ₀ := +is_noetherian.iff_rank_lt_aleph_0.1 variables {K V} /-- In a noetherian module over a division ring, all bases are indexed by a finite type. -/ noncomputable def fintype_basis_index {ι : Type*} [is_noetherian K V] (b : basis ι K V) : fintype ι := -b.fintype_index_of_dim_lt_aleph_0 (dim_lt_aleph_0 K V) +b.fintype_index_of_rank_lt_aleph_0 (rank_lt_aleph_0 K V) /-- In a noetherian module over a division ring, `basis.of_vector_space` is indexed by a finite type. -/ @@ -61,7 +61,7 @@ fintype_basis_index (basis.of_vector_space K V) if a basis is indexed by a set, that set is finite. -/ lemma finite_basis_index {ι : Type*} {s : set ι} [is_noetherian K V] (b : basis s K V) : s.finite := -b.finite_index_of_dim_lt_aleph_0 (dim_lt_aleph_0 K V) +b.finite_index_of_rank_lt_aleph_0 (rank_lt_aleph_0 K V) variables (K V) @@ -103,8 +103,8 @@ begin { introI h, exact ⟨⟨finset_basis_index K V, by { convert (finset_basis K V).span_eq, simp }⟩⟩ }, { rintros ⟨s, hs⟩, - rw [is_noetherian.iff_dim_lt_aleph_0, ← dim_top, ← hs], - exact lt_of_le_of_lt (dim_span_le _) s.finite_to_set.lt_aleph_0 } + rw [is_noetherian.iff_rank_lt_aleph_0, ← rank_top, ← hs], + exact lt_of_le_of_lt (rank_span_le _) s.finite_to_set.lt_aleph_0 } end end is_noetherian diff --git a/src/field_theory/fixed.lean b/src/field_theory/fixed.lean index f8192a0368438..115c4b0289428 100644 --- a/src/field_theory/fixed.lean +++ b/src/field_theory/fixed.lean @@ -232,10 +232,10 @@ theorem minpoly_eq_minpoly : minpoly.eq_of_irreducible_of_monic (minpoly.irreducible G F x) (minpoly.eval₂ G F x) (minpoly.monic G F x) -lemma dim_le_card : module.rank (fixed_points.subfield G F) F ≤ fintype.card G := -dim_le $ λ s hs, by simpa only [dim_fun', cardinal.mk_coe_finset, finset.coe_sort_coe, +lemma rank_le_card : module.rank (fixed_points.subfield G F) F ≤ fintype.card G := +rank_le $ λ s hs, by simpa only [rank_fun', cardinal.mk_coe_finset, finset.coe_sort_coe, cardinal.lift_nat_cast, cardinal.nat_cast_le] - using cardinal_lift_le_dim_of_linear_independent' + using cardinal_lift_le_rank_of_linear_independent' (linear_independent_smul_of_linear_independent G F hs) end fintype @@ -262,15 +262,15 @@ instance separable : is_separable (fixed_points.subfield G F) F := exact polynomial.separable_prod_X_sub_C_iff.2 (injective_of_quotient_stabilizer G x) }⟩ instance : finite_dimensional (subfield G F) F := -by { casesI nonempty_fintype G, exact is_noetherian.iff_fg.1 (is_noetherian.iff_dim_lt_aleph_0.2 $ - (dim_le_card G F).trans_lt $ cardinal.nat_lt_aleph_0 _) } +by { casesI nonempty_fintype G, exact is_noetherian.iff_fg.1 (is_noetherian.iff_rank_lt_aleph_0.2 $ + (rank_le_card G F).trans_lt $ cardinal.nat_lt_aleph_0 _) } end finite lemma finrank_le_card [fintype G] : finrank (subfield G F) F ≤ fintype.card G := begin - rw [← cardinal.nat_cast_le, finrank_eq_dim], - apply dim_le_card, + rw [← cardinal.nat_cast_le, finrank_eq_rank], + apply rank_le_card, end end fixed_points diff --git a/src/field_theory/intermediate_field.lean b/src/field_theory/intermediate_field.lean index 204b1b4415602..2306b210cb2a4 100644 --- a/src/field_theory/intermediate_field.lean +++ b/src/field_theory/intermediate_field.lean @@ -472,7 +472,7 @@ left K F L instance finite_dimensional_right [finite_dimensional K L] : finite_dimensional F L := right K F L -@[simp] lemma dim_eq_dim_subalgebra : +@[simp] lemma rank_eq_rank_subalgebra : module.rank K F.to_subalgebra = module.rank K F := rfl @[simp] lemma finrank_eq_finrank_subalgebra : diff --git a/src/field_theory/krull_topology.lean b/src/field_theory/krull_topology.lean index b24e897daa3cb..0901d5e87b26a 100644 --- a/src/field_theory/krull_topology.lean +++ b/src/field_theory/krull_topology.lean @@ -88,7 +88,7 @@ intermediate_field.fixing_subgroup '' (finite_exts K L) /-- For an field extension `L/K`, the intermediate field `K` is finite-dimensional over `K` -/ lemma intermediate_field.finite_dimensional_bot (K L : Type*) [field K] [field L] [algebra K L] : finite_dimensional K (⊥ : intermediate_field K L) := -finite_dimensional_of_dim_eq_one intermediate_field.dim_bot +finite_dimensional_of_rank_eq_one intermediate_field.rank_bot /-- This lemma says that `Gal(L/K) = L ≃ₐ[K] L` -/ lemma intermediate_field.fixing_subgroup.bot {K L : Type*} [field K] diff --git a/src/field_theory/mv_polynomial.lean b/src/field_theory/mv_polynomial.lean index d55696ff1367d..e048ed6a48c87 100644 --- a/src/field_theory/mv_polynomial.lean +++ b/src/field_theory/mv_polynomial.lean @@ -50,7 +50,7 @@ variables {σ : Type u} {K : Type u} [field K] open_locale classical -lemma dim_mv_polynomial : module.rank K (mv_polynomial σ K) = cardinal.mk (σ →₀ ℕ) := -by rw [← cardinal.lift_inj, ← (basis_monomials σ K).mk_eq_dim] +lemma rank_mv_polynomial : module.rank K (mv_polynomial σ K) = cardinal.mk (σ →₀ ℕ) := +by rw [← cardinal.lift_inj, ← (basis_monomials σ K).mk_eq_rank] end mv_polynomial diff --git a/src/field_theory/tower.lean b/src/field_theory/tower.lean index c23085bf3ebd3..0a4321cf1f790 100644 --- a/src/field_theory/tower.lean +++ b/src/field_theory/tower.lean @@ -44,21 +44,21 @@ variables [algebra F K] [module K A] [module F A] [is_scalar_tower F K A] /-- Tower law: if `A` is a `K`-vector space and `K` is a field extension of `F` then `dim_F(A) = dim_F(K) * dim_K(A)`. -/ -theorem dim_mul_dim' : +theorem rank_mul_rank' : (cardinal.lift.{w} (module.rank F K) * cardinal.lift.{v} (module.rank K A)) = cardinal.lift.{v} (module.rank F A) := let b := basis.of_vector_space F K, c := basis.of_vector_space K A in -by rw [← (module.rank F K).lift_id, ← b.mk_eq_dim, - ← (module.rank K A).lift_id, ← c.mk_eq_dim, - ← lift_umax.{w v}, ← (b.smul c).mk_eq_dim, mk_prod, lift_mul, +by rw [← (module.rank F K).lift_id, ← b.mk_eq_rank, + ← (module.rank K A).lift_id, ← c.mk_eq_rank, + ← lift_umax.{w v}, ← (b.smul c).mk_eq_rank, mk_prod, lift_mul, lift_lift, lift_lift, lift_lift, lift_lift, lift_umax] /-- Tower law: if `A` is a `K`-vector space and `K` is a field extension of `F` then `dim_F(A) = dim_F(K) * dim_K(A)`. -/ -theorem dim_mul_dim (F : Type u) (K A : Type v) [field F] [field K] [add_comm_group A] +theorem rank_mul_rank (F : Type u) (K A : Type v) [field F] [field K] [add_comm_group A] [algebra F K] [module K A] [module F A] [is_scalar_tower F K A] : module.rank F K * module.rank K A = module.rank F A := -by convert dim_mul_dim' F K A; rw lift_id +by convert rank_mul_rank' F K A; rw lift_id namespace finite_dimensional diff --git a/src/linear_algebra/affine_space/finite_dimensional.lean b/src/linear_algebra/affine_space/finite_dimensional.lean index 27419bb618522..a596e275cdac0 100644 --- a/src/linear_algebra/affine_space/finite_dimensional.lean +++ b/src/linear_algebra/affine_space/finite_dimensional.lean @@ -313,7 +313,7 @@ at most `1`. -/ def collinear (s : set P) : Prop := module.rank k (vector_span k s) ≤ 1 /-- The definition of `collinear`. -/ -lemma collinear_iff_dim_le_one (s : set P) : collinear k s ↔ module.rank k (vector_span k s) ≤ 1 := +lemma collinear_iff_rank_le_one (s : set P) : collinear k s ↔ module.rank k (vector_span k s) ≤ 1 := iff.rfl variables {k} @@ -324,8 +324,8 @@ collinear if and only if their `vector_span` has dimension at most lemma collinear_iff_finrank_le_one {s : set P} [finite_dimensional k (vector_span k s)] : collinear k s ↔ finrank k (vector_span k s) ≤ 1 := begin - have h := collinear_iff_dim_le_one k s, - rw ←finrank_eq_dim at h, + have h := collinear_iff_rank_le_one k s, + rw ←finrank_eq_rank at h, exact_mod_cast h end @@ -333,13 +333,13 @@ alias collinear_iff_finrank_le_one ↔ collinear.finrank_le_one _ /-- A subset of a collinear set is collinear. -/ lemma collinear.subset {s₁ s₂ : set P} (hs : s₁ ⊆ s₂) (h : collinear k s₂) : collinear k s₁ := -(dim_le_of_submodule (vector_span k s₁) (vector_span k s₂) (vector_span_mono k hs)).trans h +(rank_le_of_submodule (vector_span k s₁) (vector_span k s₂) (vector_span_mono k hs)).trans h /-- The `vector_span` of collinear points is finite-dimensional. -/ lemma collinear.finite_dimensional_vector_span {s : set P} (h : collinear k s) : finite_dimensional k (vector_span k s) := is_noetherian.iff_fg.1 - (is_noetherian.iff_dim_lt_aleph_0.2 (lt_of_le_of_lt h cardinal.one_lt_aleph_0)) + (is_noetherian.iff_rank_lt_aleph_0.2 (lt_of_le_of_lt h cardinal.one_lt_aleph_0)) /-- The direction of the affine span of collinear points is finite-dimensional. -/ lemma collinear.finite_dimensional_direction_affine_span {s : set P} (h : collinear k s) : @@ -351,7 +351,7 @@ variables (k P) /-- The empty set is collinear. -/ lemma collinear_empty : collinear k (∅ : set P) := begin - rw [collinear_iff_dim_le_one, vector_span_empty], + rw [collinear_iff_rank_le_one, vector_span_empty], simp end @@ -360,7 +360,7 @@ variables {P} /-- A single point is collinear. -/ lemma collinear_singleton (p : P) : collinear k ({p} : set P) := begin - rw [collinear_iff_dim_le_one, vector_span_singleton], + rw [collinear_iff_rank_le_one, vector_span_singleton], simp end @@ -372,7 +372,7 @@ vector, added to `p₀`. -/ lemma collinear_iff_of_mem {s : set P} {p₀ : P} (h : p₀ ∈ s) : collinear k s ↔ ∃ v : V, ∀ p ∈ s, ∃ r : k, p = r • v +ᵥ p₀ := begin - simp_rw [collinear_iff_dim_le_one, dim_submodule_le_one_iff', submodule.le_span_singleton_iff], + simp_rw [collinear_iff_rank_le_one, rank_submodule_le_one_iff', submodule.le_span_singleton_iff], split, { rintro ⟨v₀, hv⟩, use v₀, @@ -594,7 +594,7 @@ variables {k} lemma coplanar.finite_dimensional_vector_span {s : set P} (h : coplanar k s) : finite_dimensional k (vector_span k s) := begin - refine is_noetherian.iff_fg.1 (is_noetherian.iff_dim_lt_aleph_0.2 (lt_of_le_of_lt h _)), + refine is_noetherian.iff_fg.1 (is_noetherian.iff_rank_lt_aleph_0.2 (lt_of_le_of_lt h _)), simp, end @@ -609,7 +609,7 @@ lemma coplanar_iff_finrank_le_two {s : set P} [finite_dimensional k (vector_span coplanar k s ↔ finrank k (vector_span k s) ≤ 2 := begin have h : coplanar k s ↔ module.rank k (vector_span k s) ≤ 2 := iff.rfl, - rw ←finrank_eq_dim at h, + rw ←finrank_eq_rank at h, exact_mod_cast h end @@ -617,7 +617,7 @@ alias coplanar_iff_finrank_le_two ↔ coplanar.finrank_le_two _ /-- A subset of a coplanar set is coplanar. -/ lemma coplanar.subset {s₁ s₂ : set P} (hs : s₁ ⊆ s₂) (h : coplanar k s₂) : coplanar k s₁ := -(dim_le_of_submodule (vector_span k s₁) (vector_span k s₂) (vector_span_mono k hs)).trans h +(rank_le_of_submodule (vector_span k s₁) (vector_span k s₂) (vector_span_mono k hs)).trans h /-- Collinear points are coplanar. -/ lemma collinear.coplanar {s : set P} (h : collinear k s) : coplanar k s := @@ -681,7 +681,7 @@ begin convert rfl; simp }, { rw [affine_span_coe, direction_affine_span_insert hp₀, add_comm], - refine (submodule.dim_add_le_dim_add_dim _ _).trans (add_le_add_right _ _), + refine (submodule.rank_add_le_rank_add_rank _ _).trans (add_le_add_right _ _), refine finrank_le_one ⟨p -ᵥ p₀, submodule.mem_span_singleton_self _⟩ (λ v, _), have h := v.property, rw submodule.mem_span_singleton at h, diff --git a/src/linear_algebra/bilinear_form.lean b/src/linear_algebra/bilinear_form.lean index a76ba207c864f..aed04531ab05f 100644 --- a/src/linear_algebra/bilinear_form.lean +++ b/src/linear_algebra/bilinear_form.lean @@ -1204,7 +1204,7 @@ begin exact hx₂ n hn }, refine is_compl.of_eq this (eq_top_of_finrank_eq $ (submodule.finrank_le _).antisymm _), conv_rhs { rw ← add_zero (finrank K _) }, - rw [← finrank_bot K V, ← this, submodule.dim_sup_add_dim_inf_eq, + rw [← finrank_bot K V, ← this, submodule.rank_sup_add_rank_inf_eq, finrank_add_finrank_orthogonal b₁], exact le_self_add, end diff --git a/src/linear_algebra/dimension.lean b/src/linear_algebra/dimension.lean index 894f946ba1206..61da11984d18c 100644 --- a/src/linear_algebra/dimension.lean +++ b/src/linear_algebra/dimension.lean @@ -23,9 +23,9 @@ import set_theory.cardinal.cofinality ## Main statements -* `linear_map.dim_le_of_injective`: the source of an injective linear map has dimension +* `linear_map.rank_le_of_injective`: the source of an injective linear map has dimension at most that of the target. -* `linear_map.dim_le_of_surjective`: the target of a surjective linear map has dimension +* `linear_map.rank_le_of_surjective`: the target of a surjective linear map has dimension at most that of that source. * `basis_fintype_of_finite_spans`: the existence of a finite spanning set implies that any basis is finite. @@ -58,15 +58,15 @@ For modules over rings with invariant basis number For vector spaces (i.e. modules over a field), we have -* `dim_quotient_add_dim`: if `V₁` is a submodule of `V`, then +* `rank_quotient_add_rank`: if `V₁` is a submodule of `V`, then `module.rank (V/V₁) + module.rank V₁ = module.rank V`. -* `dim_range_add_dim_ker`: the rank-nullity theorem. +* `rank_range_add_rank_ker`: the rank-nullity theorem. ## Implementation notes -There is a naming discrepancy: most of the theorem names refer to `dim`, +There is a naming discrepancy: most of the theorem names refer to `rank`, even though the definition is of `module.rank`. -This reflects that `module.rank` was originally called `dim`, and only defined for vector spaces. +This reflects that `module.rank` was originally called `rank`, and only defined for vector spaces. Many theorems in this file are not universe-generic when they relate dimensions in different universes. They should be as general as they can be without @@ -117,7 +117,7 @@ variables {M : Type v} [add_comm_group M] [module R M] variables {M' : Type v'} [add_comm_group M'] [module R M'] variables {M₁ : Type v} [add_comm_group M₁] [module R M₁] -theorem linear_map.lift_dim_le_of_injective (f : M →ₗ[R] M') (i : injective f) : +theorem linear_map.lift_rank_le_of_injective (f : M →ₗ[R] M') (i : injective f) : cardinal.lift.{v'} (module.rank R M) ≤ cardinal.lift.{v} (module.rank R M') := begin dsimp [module.rank], @@ -129,11 +129,11 @@ begin exact (li.map' _ $ linear_map.ker_eq_bot.mpr i).image, end -theorem linear_map.dim_le_of_injective (f : M →ₗ[R] M₁) (i : injective f) : +theorem linear_map.rank_le_of_injective (f : M →ₗ[R] M₁) (i : injective f) : module.rank R M ≤ module.rank R M₁ := -cardinal.lift_le.1 (f.lift_dim_le_of_injective i) +cardinal.lift_le.1 (f.lift_rank_le_of_injective i) -theorem dim_le {n : ℕ} +theorem rank_le {n : ℕ} (H : ∀ s : finset M, linear_independent R (λ i : s, (i : M)) → s.card ≤ n) : module.rank R M ≤ n := begin @@ -143,7 +143,7 @@ begin exact linear_independent_bounded_of_finset_linear_independent_bounded H _ li, end -lemma lift_dim_range_le (f : M →ₗ[R] M') : +lemma lift_rank_range_le (f : M →ₗ[R] M') : cardinal.lift.{v} (module.rank R f.range) ≤ cardinal.lift.{v'} (module.rank R M) := begin dsimp [module.rank], @@ -159,82 +159,83 @@ begin { exact (cardinal.lift_mk_eq'.mpr ⟨equiv.set.range_splitting_image_equiv f s⟩).ge, }, end -lemma dim_range_le (f : M →ₗ[R] M₁) : module.rank R f.range ≤ module.rank R M := -by simpa using lift_dim_range_le f +lemma rank_range_le (f : M →ₗ[R] M₁) : module.rank R f.range ≤ module.rank R M := +by simpa using lift_rank_range_le f -lemma lift_dim_map_le (f : M →ₗ[R] M') (p : submodule R M) : +lemma lift_rank_map_le (f : M →ₗ[R] M') (p : submodule R M) : cardinal.lift.{v} (module.rank R (p.map f)) ≤ cardinal.lift.{v'} (module.rank R p) := begin - have h := lift_dim_range_le (f.comp (submodule.subtype p)), + have h := lift_rank_range_le (f.comp (submodule.subtype p)), rwa [linear_map.range_comp, range_subtype] at h, end -lemma dim_map_le (f : M →ₗ[R] M₁) (p : submodule R M) : module.rank R (p.map f) ≤ module.rank R p := -by simpa using lift_dim_map_le f p +lemma rank_map_le (f : M →ₗ[R] M₁) (p : submodule R M) : + module.rank R (p.map f) ≤ module.rank R p := +by simpa using lift_rank_map_le f p -lemma dim_le_of_submodule (s t : submodule R M) (h : s ≤ t) : +lemma rank_le_of_submodule (s t : submodule R M) (h : s ≤ t) : module.rank R s ≤ module.rank R t := -(of_le h).dim_le_of_injective $ assume ⟨x, hx⟩ ⟨y, hy⟩ eq, +(of_le h).rank_le_of_injective $ assume ⟨x, hx⟩ ⟨y, hy⟩ eq, subtype.eq $ show x = y, from subtype.ext_iff_val.1 eq /-- Two linearly equivalent vector spaces have the same dimension, a version with different universes. -/ -theorem linear_equiv.lift_dim_eq (f : M ≃ₗ[R] M') : +theorem linear_equiv.lift_rank_eq (f : M ≃ₗ[R] M') : cardinal.lift.{v'} (module.rank R M) = cardinal.lift.{v} (module.rank R M') := begin apply le_antisymm, - { exact f.to_linear_map.lift_dim_le_of_injective f.injective, }, - { exact f.symm.to_linear_map.lift_dim_le_of_injective f.symm.injective, }, + { exact f.to_linear_map.lift_rank_le_of_injective f.injective, }, + { exact f.symm.to_linear_map.lift_rank_le_of_injective f.symm.injective, }, end /-- Two linearly equivalent vector spaces have the same dimension. -/ -theorem linear_equiv.dim_eq (f : M ≃ₗ[R] M₁) : +theorem linear_equiv.rank_eq (f : M ≃ₗ[R] M₁) : module.rank R M = module.rank R M₁ := -cardinal.lift_inj.1 f.lift_dim_eq +cardinal.lift_inj.1 f.lift_rank_eq -lemma dim_eq_of_injective (f : M →ₗ[R] M₁) (h : injective f) : +lemma rank_eq_of_injective (f : M →ₗ[R] M₁) (h : injective f) : module.rank R M = module.rank R f.range := -(linear_equiv.of_injective f h).dim_eq +(linear_equiv.of_injective f h).rank_eq /-- Pushforwards of submodules along a `linear_equiv` have the same dimension. -/ -lemma linear_equiv.dim_map_eq (f : M ≃ₗ[R] M₁) (p : submodule R M) : +lemma linear_equiv.rank_map_eq (f : M ≃ₗ[R] M₁) (p : submodule R M) : module.rank R (p.map (f : M →ₗ[R] M₁)) = module.rank R p := -(f.submodule_map p).dim_eq.symm +(f.submodule_map p).rank_eq.symm variables (R M) -@[simp] lemma dim_top : module.rank R (⊤ : submodule R M) = module.rank R M := +@[simp] lemma rank_top : module.rank R (⊤ : submodule R M) = module.rank R M := begin have : (⊤ : submodule R M) ≃ₗ[R] M := linear_equiv.of_top ⊤ rfl, - rw this.dim_eq, + rw this.rank_eq, end variables {R M} -lemma dim_range_of_surjective (f : M →ₗ[R] M') (h : surjective f) : +lemma rank_range_of_surjective (f : M →ₗ[R] M') (h : surjective f) : module.rank R f.range = module.rank R M' := -by rw [linear_map.range_eq_top.2 h, dim_top] +by rw [linear_map.range_eq_top.2 h, rank_top] -lemma dim_submodule_le (s : submodule R M) : module.rank R s ≤ module.rank R M := +lemma rank_submodule_le (s : submodule R M) : module.rank R s ≤ module.rank R M := begin - rw ←dim_top R M, - exact dim_le_of_submodule _ _ le_top, + rw ←rank_top R M, + exact rank_le_of_submodule _ _ le_top, end -lemma linear_map.dim_le_of_surjective (f : M →ₗ[R] M₁) (h : surjective f) : +lemma linear_map.rank_le_of_surjective (f : M →ₗ[R] M₁) (h : surjective f) : module.rank R M₁ ≤ module.rank R M := begin - rw ←dim_range_of_surjective f h, - apply dim_range_le, + rw ←rank_range_of_surjective f h, + apply rank_range_le, end -theorem dim_quotient_le (p : submodule R M) : +theorem rank_quotient_le (p : submodule R M) : module.rank R (M ⧸ p) ≤ module.rank R M := -(mkq p).dim_le_of_surjective (surjective_quot_mk _) +(mkq p).rank_le_of_surjective (surjective_quot_mk _) variables [nontrivial R] -lemma {m} cardinal_lift_le_dim_of_linear_independent +lemma {m} cardinal_lift_le_rank_of_linear_independent {ι : Type w} {v : ι → M} (hv : linear_independent R v) : cardinal.lift.{max v m} (#ι) ≤ cardinal.lift.{max w m} (module.rank R M) := begin @@ -248,24 +249,24 @@ begin exact le_rfl, }, end -lemma cardinal_lift_le_dim_of_linear_independent' +lemma cardinal_lift_le_rank_of_linear_independent' {ι : Type w} {v : ι → M} (hv : linear_independent R v) : cardinal.lift.{v} (#ι) ≤ cardinal.lift.{w} (module.rank R M) := -cardinal_lift_le_dim_of_linear_independent.{u v w 0} hv +cardinal_lift_le_rank_of_linear_independent.{u v w 0} hv -lemma cardinal_le_dim_of_linear_independent +lemma cardinal_le_rank_of_linear_independent {ι : Type v} {v : ι → M} (hv : linear_independent R v) : #ι ≤ module.rank R M := -by simpa using cardinal_lift_le_dim_of_linear_independent hv +by simpa using cardinal_lift_le_rank_of_linear_independent hv -lemma cardinal_le_dim_of_linear_independent' +lemma cardinal_le_rank_of_linear_independent' {s : set M} (hs : linear_independent R (λ x, x : s → M)) : #s ≤ module.rank R M := -cardinal_le_dim_of_linear_independent hs +cardinal_le_rank_of_linear_independent hs variables (R M) -@[simp] lemma dim_punit : module.rank R punit = 0 := +@[simp] lemma rank_punit : module.rank R punit = 0 := begin apply le_bot_iff.mp, rw module.rank, @@ -279,17 +280,17 @@ begin simpa using linear_independent.ne_zero (⟨a, ha⟩ : s) li, end -@[simp] lemma dim_bot : module.rank R (⊥ : submodule R M) = 0 := +@[simp] lemma rank_bot : module.rank R (⊥ : submodule R M) = 0 := begin have : (⊥ : submodule R M) ≃ₗ[R] punit := bot_equiv_punit, - rw [this.dim_eq, dim_punit], + rw [this.rank_eq, rank_punit], end variables {R M} -lemma exists_mem_ne_zero_of_dim_pos {s : submodule R M} (h : 0 < module.rank R s) : +lemma exists_mem_ne_zero_of_rank_pos {s : submodule R M} (h : 0 < module.rank R s) : ∃ b : M, b ∈ s ∧ b ≠ 0 := -exists_mem_ne_zero_of_ne_bot $ assume eq, by rw [eq, dim_bot] at h; exact lt_irrefl _ h +exists_mem_ne_zero_of_ne_bot $ assume eq, by rw [eq, rank_bot] at h; exact lt_irrefl _ h /-- A linearly-independent family of vectors in a module over a non-trivial ring must be finite if the module is Noetherian. -/ @@ -460,7 +461,7 @@ begin choose v hvV hv using hI, have : linear_independent R v, { exact (hV.comp subtype.coe_injective).linear_independent _ hvV hv }, - exact cardinal_lift_le_dim_of_linear_independent' this + exact cardinal_lift_le_rank_of_linear_independent' this end end @@ -470,7 +471,7 @@ section rank_zero variables {R : Type u} {M : Type v} variables [ring R] [add_comm_group M] [module R M] -@[simp] lemma dim_subsingleton [subsingleton R] : module.rank R M = 1 := +@[simp] lemma rank_subsingleton [subsingleton R] : module.rank R M = 1 := begin haveI := module.subsingleton R M, haveI : nonempty {s : set M // linear_independent R (coe : s → M)}, @@ -489,43 +490,43 @@ end variables [no_zero_smul_divisors R M] -lemma dim_pos [nontrivial M] : 0 < module.rank R M := +lemma rank_pos [nontrivial M] : 0 < module.rank R M := begin obtain ⟨x, hx⟩ := exists_ne (0 : M), suffices : 1 ≤ module.rank R M, { exact zero_lt_one.trans_le this }, letI := module.nontrivial R M, suffices : linear_independent R (λ (y : ({x} : set M)), ↑y), - { simpa using (cardinal_le_dim_of_linear_independent this), }, + { simpa using (cardinal_le_rank_of_linear_independent this), }, exact linear_independent_singleton hx end variables [nontrivial R] -lemma dim_zero_iff_forall_zero : module.rank R M = 0 ↔ ∀ x : M, x = 0 := +lemma rank_zero_iff_forall_zero : module.rank R M = 0 ↔ ∀ x : M, x = 0 := begin refine ⟨λ h, _, λ h, _⟩, { contrapose! h, obtain ⟨x, hx⟩ := h, letI : nontrivial M := nontrivial_of_ne _ _ hx, - exact dim_pos.ne' }, + exact rank_pos.ne' }, { have : (⊤ : submodule R M) = ⊥, { ext x, simp [h x] }, - rw [←dim_top, this, dim_bot] } + rw [←rank_top, this, rank_bot] } end -/-- See `dim_subsingleton` for the reason that `nontrivial R` is needed. -/ -lemma dim_zero_iff : module.rank R M = 0 ↔ subsingleton M := -dim_zero_iff_forall_zero.trans (subsingleton_iff_forall_eq 0).symm +/-- See `rank_subsingleton` for the reason that `nontrivial R` is needed. -/ +lemma rank_zero_iff : module.rank R M = 0 ↔ subsingleton M := +rank_zero_iff_forall_zero.trans (subsingleton_iff_forall_eq 0).symm -lemma dim_pos_iff_exists_ne_zero : 0 < module.rank R M ↔ ∃ x : M, x ≠ 0 := +lemma rank_pos_iff_exists_ne_zero : 0 < module.rank R M ↔ ∃ x : M, x ≠ 0 := begin rw ←not_iff_not, - simpa using dim_zero_iff_forall_zero + simpa using rank_zero_iff_forall_zero end -lemma dim_pos_iff_nontrivial : 0 < module.rank R M ↔ nontrivial M := -dim_pos_iff_exists_ne_zero.trans (nontrivial_iff_exists_ne 0).symm +lemma rank_pos_iff_nontrivial : 0 < module.rank R M ↔ nontrivial M := +rank_pos_iff_exists_ne_zero.trans (nontrivial_iff_exists_ne 0).symm end rank_zero @@ -805,7 +806,7 @@ begin exact infinite_basis_le_maximal_linear_independent b v i m, } end -theorem basis.mk_eq_dim'' {ι : Type v} (v : basis ι R M) : +theorem basis.mk_eq_rank'' {ι : Type v} (v : basis ι R M) : #ι = module.rank R M := begin haveI := nontrivial_of_invariant_basis_number R, @@ -821,24 +822,24 @@ begin apply linear_independent_le_basis v _ li, }, end -theorem basis.mk_range_eq_dim (v : basis ι R M) : +theorem basis.mk_range_eq_rank (v : basis ι R M) : #(range v) = module.rank R M := -v.reindex_range.mk_eq_dim'' +v.reindex_range.mk_eq_rank'' /-- If a vector space has a finite basis, then its dimension (seen as a cardinal) is equal to the cardinality of the basis. -/ -lemma dim_eq_card_basis {ι : Type w} [fintype ι] (h : basis ι R M) : +lemma rank_eq_card_basis {ι : Type w} [fintype ι] (h : basis ι R M) : module.rank R M = fintype.card ι := by {haveI := nontrivial_of_invariant_basis_number R, - rw [←h.mk_range_eq_dim, cardinal.mk_fintype, set.card_range_of_injective h.injective] } + rw [←h.mk_range_eq_rank, cardinal.mk_fintype, set.card_range_of_injective h.injective] } lemma basis.card_le_card_of_linear_independent {ι : Type*} [fintype ι] (b : basis ι R M) {ι' : Type*} [fintype ι'] {v : ι' → M} (hv : linear_independent R v) : fintype.card ι' ≤ fintype.card ι := begin letI := nontrivial_of_invariant_basis_number R, - simpa [dim_eq_card_basis b, cardinal.mk_fintype] using - cardinal_lift_le_dim_of_linear_independent' hv + simpa [rank_eq_card_basis b, cardinal.mk_fintype] using + cardinal_lift_le_rank_of_linear_independent' hv end lemma basis.card_le_card_of_submodule (N : submodule R M) [fintype ι] (b : basis ι R M) @@ -851,49 +852,49 @@ lemma basis.card_le_card_of_le b.card_le_card_of_linear_independent (b'.linear_independent.map' (submodule.of_le hNO) (N.ker_of_le O _)) -theorem basis.mk_eq_dim (v : basis ι R M) : +theorem basis.mk_eq_rank (v : basis ι R M) : cardinal.lift.{v} (#ι) = cardinal.lift.{w} (module.rank R M) := begin haveI := nontrivial_of_invariant_basis_number R, - rw [←v.mk_range_eq_dim, cardinal.mk_range_eq_of_injective v.injective] + rw [←v.mk_range_eq_rank, cardinal.mk_range_eq_of_injective v.injective] end -theorem {m} basis.mk_eq_dim' (v : basis ι R M) : +theorem {m} basis.mk_eq_rank' (v : basis ι R M) : cardinal.lift.{max v m} (#ι) = cardinal.lift.{max w m} (module.rank R M) := -by simpa using v.mk_eq_dim +by simpa using v.mk_eq_rank /-- If a module has a finite dimension, all bases are indexed by a finite type. -/ -lemma basis.nonempty_fintype_index_of_dim_lt_aleph_0 {ι : Type*} +lemma basis.nonempty_fintype_index_of_rank_lt_aleph_0 {ι : Type*} (b : basis ι R M) (h : module.rank R M < ℵ₀) : nonempty (fintype ι) := -by rwa [← cardinal.lift_lt, ← b.mk_eq_dim, +by rwa [← cardinal.lift_lt, ← b.mk_eq_rank, -- ensure `aleph_0` has the correct universe cardinal.lift_aleph_0, ← cardinal.lift_aleph_0.{u_1 v}, cardinal.lift_lt, cardinal.lt_aleph_0_iff_fintype] at h /-- If a module has a finite dimension, all bases are indexed by a finite type. -/ -noncomputable def basis.fintype_index_of_dim_lt_aleph_0 {ι : Type*} +noncomputable def basis.fintype_index_of_rank_lt_aleph_0 {ι : Type*} (b : basis ι R M) (h : module.rank R M < ℵ₀) : fintype ι := -classical.choice (b.nonempty_fintype_index_of_dim_lt_aleph_0 h) +classical.choice (b.nonempty_fintype_index_of_rank_lt_aleph_0 h) /-- If a module has a finite dimension, all bases are indexed by a finite set. -/ -lemma basis.finite_index_of_dim_lt_aleph_0 {ι : Type*} {s : set ι} +lemma basis.finite_index_of_rank_lt_aleph_0 {ι : Type*} {s : set ι} (b : basis s R M) (h : module.rank R M < ℵ₀) : s.finite := -finite_def.2 (b.nonempty_fintype_index_of_dim_lt_aleph_0 h) +finite_def.2 (b.nonempty_fintype_index_of_rank_lt_aleph_0 h) -lemma dim_span {v : ι → M} (hv : linear_independent R v) : +lemma rank_span {v : ι → M} (hv : linear_independent R v) : module.rank R ↥(span R (range v)) = #(range v) := begin haveI := nontrivial_of_invariant_basis_number R, - rw [←cardinal.lift_inj, ← (basis.span hv).mk_eq_dim, + rw [←cardinal.lift_inj, ← (basis.span hv).mk_eq_rank, cardinal.mk_range_eq_of_injective (@linear_independent.injective ι R M v _ _ _ _ hv)] end -lemma dim_span_set {s : set M} (hs : linear_independent R (λ x, x : s → M)) : +lemma rank_span_set {s : set M} (hs : linear_independent R (λ x, x : s → M)) : module.rank R ↥(span R s) = #s := -by { rw [← @set_of_mem_eq _ s, ← subtype.range_coe_subtype], exact dim_span hs } +by { rw [← @set_of_mem_eq _ s, ← subtype.range_coe_subtype], exact rank_span hs } /-- If `N` is a submodule in a free, finitely generated module, do induction on adjoining a linear independent element to a submodule. -/ @@ -929,8 +930,8 @@ end variables (R) -@[simp] lemma dim_self : module.rank R R = 1 := -by rw [←cardinal.lift_inj, ← (basis.singleton punit R).mk_eq_dim, cardinal.mk_punit] +@[simp] lemma rank_self : module.rank R R = 1 := +by rw [←cardinal.lift_inj, ← (basis.singleton punit R).mk_eq_rank, cardinal.mk_punit] end strong_rank_condition @@ -942,58 +943,58 @@ variables [add_comm_group V₁] [module K V₁] [module.free K V₁] variables {K V} /-- Two vector spaces are isomorphic if they have the same dimension. -/ -theorem nonempty_linear_equiv_of_lift_dim_eq +theorem nonempty_linear_equiv_of_lift_rank_eq (cond : cardinal.lift.{v'} (module.rank K V) = cardinal.lift.{v} (module.rank K V')) : nonempty (V ≃ₗ[K] V') := begin obtain ⟨⟨_, B⟩⟩ := module.free.exists_basis K V, obtain ⟨⟨_, B'⟩⟩ := module.free.exists_basis K V', have : cardinal.lift.{v' v} (#_) = cardinal.lift.{v v'} (#_), - by rw [B.mk_eq_dim'', cond, B'.mk_eq_dim''], + by rw [B.mk_eq_rank'', cond, B'.mk_eq_rank''], exact (cardinal.lift_mk_eq.{v v' 0}.1 this).map (B.equiv B') end /-- Two vector spaces are isomorphic if they have the same dimension. -/ -theorem nonempty_linear_equiv_of_dim_eq +theorem nonempty_linear_equiv_of_rank_eq (cond : module.rank K V = module.rank K V₁) : nonempty (V ≃ₗ[K] V₁) := -nonempty_linear_equiv_of_lift_dim_eq $ congr_arg _ cond +nonempty_linear_equiv_of_lift_rank_eq $ congr_arg _ cond section variables (V V' V₁) /-- Two vector spaces are isomorphic if they have the same dimension. -/ -def linear_equiv.of_lift_dim_eq +def linear_equiv.of_lift_rank_eq (cond : cardinal.lift.{v'} (module.rank K V) = cardinal.lift.{v} (module.rank K V')) : V ≃ₗ[K] V' := -classical.choice (nonempty_linear_equiv_of_lift_dim_eq cond) +classical.choice (nonempty_linear_equiv_of_lift_rank_eq cond) /-- Two vector spaces are isomorphic if they have the same dimension. -/ -def linear_equiv.of_dim_eq (cond : module.rank K V = module.rank K V₁) : V ≃ₗ[K] V₁ := -classical.choice (nonempty_linear_equiv_of_dim_eq cond) +def linear_equiv.of_rank_eq (cond : module.rank K V = module.rank K V₁) : V ≃ₗ[K] V₁ := +classical.choice (nonempty_linear_equiv_of_rank_eq cond) end /-- Two vector spaces are isomorphic if and only if they have the same dimension. -/ -theorem linear_equiv.nonempty_equiv_iff_lift_dim_eq : +theorem linear_equiv.nonempty_equiv_iff_lift_rank_eq : nonempty (V ≃ₗ[K] V') ↔ cardinal.lift.{v'} (module.rank K V) = cardinal.lift.{v} (module.rank K V') := -⟨λ ⟨h⟩, linear_equiv.lift_dim_eq h, λ h, nonempty_linear_equiv_of_lift_dim_eq h⟩ +⟨λ ⟨h⟩, linear_equiv.lift_rank_eq h, λ h, nonempty_linear_equiv_of_lift_rank_eq h⟩ /-- Two vector spaces are isomorphic if and only if they have the same dimension. -/ -theorem linear_equiv.nonempty_equiv_iff_dim_eq : +theorem linear_equiv.nonempty_equiv_iff_rank_eq : nonempty (V ≃ₗ[K] V₁) ↔ module.rank K V = module.rank K V₁ := -⟨λ ⟨h⟩, linear_equiv.dim_eq h, λ h, nonempty_linear_equiv_of_dim_eq h⟩ +⟨λ ⟨h⟩, linear_equiv.rank_eq h, λ h, nonempty_linear_equiv_of_rank_eq h⟩ -theorem dim_prod : module.rank K (V × V₁) = module.rank K V + module.rank K V₁ := +theorem rank_prod : module.rank K (V × V₁) = module.rank K V + module.rank K V₁ := begin obtain ⟨⟨_, b⟩⟩ := module.free.exists_basis K V, obtain ⟨⟨_, c⟩⟩ := module.free.exists_basis K V₁, rw [← cardinal.lift_inj, - ← (basis.prod b c).mk_eq_dim, + ← (basis.prod b c).mk_eq_rank, cardinal.lift_add, ← cardinal.mk_ulift, - ← b.mk_eq_dim, ← c.mk_eq_dim, + ← b.mk_eq_rank, ← c.mk_eq_rank, ← cardinal.mk_ulift, ← cardinal.mk_ulift, cardinal.add_def (ulift _)], exact cardinal.lift_inj.1 (cardinal.lift_mk_eq.2 @@ -1005,42 +1006,42 @@ variables [∀i, add_comm_group (φ i)] [∀i, module K (φ i)] [∀i, module.fr open linear_map -lemma dim_pi [finite η] : +lemma rank_pi [finite η] : module.rank K (Πi, φ i) = cardinal.sum (λi, module.rank K (φ i)) := begin haveI := nontrivial_of_invariant_basis_number K, casesI nonempty_fintype η, let b := λ i, (module.free.exists_basis K (φ i)).some.2, let this : basis (Σ j, _) K (Π j, φ j) := pi.basis b, - rw [← cardinal.lift_inj, ← this.mk_eq_dim], - simp_rw [cardinal.mk_sigma, cardinal.lift_sum, ←(b _).mk_range_eq_dim, + rw [← cardinal.lift_inj, ← this.mk_eq_rank], + simp_rw [cardinal.mk_sigma, cardinal.lift_sum, ←(b _).mk_range_eq_rank, cardinal.mk_range_eq _ (b _).injective], end variable [fintype η] -lemma dim_fun {V η : Type u} [fintype η] [add_comm_group V] [module K V] +lemma rank_fun {V η : Type u} [fintype η] [add_comm_group V] [module K V] [module.free K V] : module.rank K (η → V) = fintype.card η * module.rank K V := -by rw [dim_pi, cardinal.sum_const', cardinal.mk_fintype] +by rw [rank_pi, cardinal.sum_const', cardinal.mk_fintype] -lemma dim_fun_eq_lift_mul : +lemma rank_fun_eq_lift_mul : module.rank K (η → V) = (fintype.card η : cardinal.{max u₁' v}) * cardinal.lift.{u₁'} (module.rank K V) := -by rw [dim_pi, cardinal.sum_const, cardinal.mk_fintype, cardinal.lift_nat_cast] +by rw [rank_pi, cardinal.sum_const, cardinal.mk_fintype, cardinal.lift_nat_cast] -lemma dim_fun' : module.rank K (η → K) = fintype.card η := -by rw [dim_fun_eq_lift_mul, dim_self, cardinal.lift_one, mul_one, cardinal.nat_cast_inj] +lemma rank_fun' : module.rank K (η → K) = fintype.card η := +by rw [rank_fun_eq_lift_mul, rank_self, cardinal.lift_one, mul_one, cardinal.nat_cast_inj] -lemma dim_fin_fun (n : ℕ) : module.rank K (fin n → K) = n := -by simp [dim_fun'] +lemma rank_fin_fun (n : ℕ) : module.rank K (fin n → K) = n := +by simp [rank_fun'] end fintype -lemma finsupp.dim_eq {ι : Type v} : module.rank K (ι →₀ V) = #ι * module.rank K V := +lemma finsupp.rank_eq {ι : Type v} : module.rank K (ι →₀ V) = #ι * module.rank K V := begin obtain ⟨⟨_, bs⟩⟩ := module.free.exists_basis K V, - rw [← bs.mk_eq_dim'', ← (finsupp.basis (λa:ι, bs)).mk_eq_dim'', + rw [← bs.mk_eq_rank'', ← (finsupp.basis (λa:ι, bs)).mk_eq_rank'', cardinal.mk_sigma, cardinal.sum_const'] end @@ -1054,9 +1055,9 @@ begin by simp, have hn := cardinal.lift_inj.{v u}.2 hn, rw this at hn, - rw ←@dim_fin_fun K _ _ n at hn, + rw ←@rank_fin_fun K _ _ n at hn, haveI : module.free K (fin n → K) := module.free.pi _ _, - exact classical.choice (nonempty_linear_equiv_of_lift_dim_eq hn), + exact classical.choice (nonempty_linear_equiv_of_lift_rank_eq hn), end end free @@ -1069,49 +1070,49 @@ variables [add_comm_group V₁] [module K V₁] variables {K V} /-- If a vector space has a finite dimension, the index set of `basis.of_vector_space` is finite. -/ -lemma basis.finite_of_vector_space_index_of_dim_lt_aleph_0 (h : module.rank K V < ℵ₀) : +lemma basis.finite_of_vector_space_index_of_rank_lt_aleph_0 (h : module.rank K V < ℵ₀) : (basis.of_vector_space_index K V).finite := -finite_def.2 $ (basis.of_vector_space K V).nonempty_fintype_index_of_dim_lt_aleph_0 h +finite_def.2 $ (basis.of_vector_space K V).nonempty_fintype_index_of_rank_lt_aleph_0 h -- TODO how far can we generalise this? -- When `s` is finite, we could prove this for any ring satisfying the strong rank condition -- using `linear_independent_le_span'` -lemma dim_span_le (s : set V) : module.rank K (span K s) ≤ #s := +lemma rank_span_le (s : set V) : module.rank K (span K s) ≤ #s := begin obtain ⟨b, hb, hsab, hlib⟩ := exists_linear_independent K s, convert cardinal.mk_le_mk_of_subset hb, - rw [← hsab, dim_span_set hlib] + rw [← hsab, rank_span_set hlib] end -lemma dim_span_of_finset (s : finset V) : +lemma rank_span_of_finset (s : finset V) : module.rank K (span K (↑s : set V)) < ℵ₀ := -calc module.rank K (span K (↑s : set V)) ≤ #(↑s : set V) : dim_span_le ↑s +calc module.rank K (span K (↑s : set V)) ≤ #(↑s : set V) : rank_span_le ↑s ... = s.card : by rw [finset.coe_sort_coe, cardinal.mk_coe_finset] ... < ℵ₀ : cardinal.nat_lt_aleph_0 _ -theorem dim_quotient_add_dim (p : submodule K V) : +theorem rank_quotient_add_rank (p : submodule K V) : module.rank K (V ⧸ p) + module.rank K p = module.rank K V := -by classical; exact let ⟨f⟩ := quotient_prod_linear_equiv p in dim_prod.symm.trans f.dim_eq +by classical; exact let ⟨f⟩ := quotient_prod_linear_equiv p in rank_prod.symm.trans f.rank_eq /-- rank-nullity theorem -/ -theorem dim_range_add_dim_ker (f : V →ₗ[K] V₁) : +theorem rank_range_add_rank_ker (f : V →ₗ[K] V₁) : module.rank K f.range + module.rank K f.ker = module.rank K V := begin haveI := λ (p : submodule K V), classical.dec_eq (V ⧸ p), - rw [← f.quot_ker_equiv_range.dim_eq, dim_quotient_add_dim] + rw [← f.quot_ker_equiv_range.rank_eq, rank_quotient_add_rank] end -lemma dim_eq_of_surjective (f : V →ₗ[K] V₁) (h : surjective f) : +lemma rank_eq_of_surjective (f : V →ₗ[K] V₁) (h : surjective f) : module.rank K V = module.rank K V₁ + module.rank K f.ker := -by rw [← dim_range_add_dim_ker f, ← dim_range_of_surjective f h] +by rw [← rank_range_add_rank_ker f, ← rank_range_of_surjective f h] section variables [add_comm_group V₂] [module K V₂] variables [add_comm_group V₃] [module K V₃] open linear_map -/-- This is mostly an auxiliary lemma for `dim_sup_add_dim_inf_eq`. -/ -lemma dim_add_dim_split +/-- This is mostly an auxiliary lemma for `rank_sup_add_rank_inf_eq`. -/ +lemma rank_add_rank_split (db : V₂ →ₗ[K] V) (eb : V₃ →ₗ[K] V) (cd : V₁ →ₗ[K] V₂) (ce : V₁ →ₗ[K] V₃) (hde : ⊤ ≤ db.range ⊔ eb.range) (hgd : ker cd = ⊥) @@ -1121,9 +1122,9 @@ lemma dim_add_dim_split have hf : surjective (coprod db eb), by rwa [←range_eq_top, range_coprod, eq_top_iff], begin - conv {to_rhs, rw [← dim_prod, dim_eq_of_surjective _ hf] }, + conv {to_rhs, rw [← rank_prod, rank_eq_of_surjective _ hf] }, congr' 1, - apply linear_equiv.dim_eq, + apply linear_equiv.rank_eq, refine linear_equiv.of_bijective _ ⟨_, _⟩, { refine cod_restrict _ (prod cd (- ce)) _, { assume c, @@ -1143,10 +1144,11 @@ begin rw [h₂, _root_.neg_neg] } end -lemma dim_sup_add_dim_inf_eq (s t : submodule K V) : +lemma rank_sup_add_rank_inf_eq (s t : submodule K V) : module.rank K (s ⊔ t : submodule K V) + module.rank K (s ⊓ t : submodule K V) = module.rank K s + module.rank K t := -dim_add_dim_split (of_le le_sup_left) (of_le le_sup_right) (of_le inf_le_left) (of_le inf_le_right) +rank_add_rank_split + (of_le le_sup_left) (of_le le_sup_right) (of_le inf_le_left) (of_le inf_le_right) begin rw [← map_le_map_iff' (ker_subtype $ s ⊔ t), submodule.map_sup, submodule.map_top, ← linear_map.range_comp, ← linear_map.range_comp, subtype_comp_of_le, subtype_comp_of_le, @@ -1161,9 +1163,9 @@ dim_add_dim_split (of_le le_sup_left) (of_le le_sup_right) (of_le inf_le_left) ( exact ⟨⟨b₁, hb₁, hb₂⟩, rfl, rfl⟩ end -lemma dim_add_le_dim_add_dim (s t : submodule K V) : +lemma rank_add_le_rank_add_rank (s t : submodule K V) : module.rank K (s ⊔ t : submodule K V) ≤ module.rank K s + module.rank K t := -by { rw [← dim_sup_add_dim_inf_eq], exact self_le_add_right _ _ } +by { rw [← rank_sup_add_rank_inf_eq], exact self_le_add_right _ _ } end @@ -1177,36 +1179,36 @@ variables [add_comm_group V'] [module K V'] See also `finite_dimensional.fin_basis`. -/ -def basis.of_dim_eq_zero {ι : Type*} [is_empty ι] (hV : module.rank K V = 0) : +def basis.of_rank_eq_zero {ι : Type*} [is_empty ι] (hV : module.rank K V = 0) : basis ι K V := begin - haveI : subsingleton V := dim_zero_iff.1 hV, + haveI : subsingleton V := rank_zero_iff.1 hV, exact basis.empty _ end -@[simp] lemma basis.of_dim_eq_zero_apply {ι : Type*} [is_empty ι] +@[simp] lemma basis.of_rank_eq_zero_apply {ι : Type*} [is_empty ι] (hV : module.rank K V = 0) (i : ι) : - basis.of_dim_eq_zero hV i = 0 := + basis.of_rank_eq_zero hV i = 0 := rfl -lemma le_dim_iff_exists_linear_independent {c : cardinal} : +lemma le_rank_iff_exists_linear_independent {c : cardinal} : c ≤ module.rank K V ↔ ∃ s : set V, #s = c ∧ linear_independent K (coe : s → V) := begin split, { intro h, let t := basis.of_vector_space K V, - rw [← t.mk_eq_dim'', cardinal.le_mk_iff_exists_subset] at h, + rw [← t.mk_eq_rank'', cardinal.le_mk_iff_exists_subset] at h, rcases h with ⟨s, hst, hsc⟩, exact ⟨s, hsc, (of_vector_space_index.linear_independent K V).mono hst⟩ }, { rintro ⟨s, rfl, si⟩, - exact cardinal_le_dim_of_linear_independent si } + exact cardinal_le_rank_of_linear_independent si } end -lemma le_dim_iff_exists_linear_independent_finset {n : ℕ} : +lemma le_rank_iff_exists_linear_independent_finset {n : ℕ} : ↑n ≤ module.rank K V ↔ ∃ s : finset V, s.card = n ∧ linear_independent K (coe : (s : set V) → V) := begin - simp only [le_dim_iff_exists_linear_independent, cardinal.mk_set_eq_nat_iff_finset], + simp only [le_rank_iff_exists_linear_independent, cardinal.mk_set_eq_nat_iff_finset], split, { rintro ⟨s, ⟨t, rfl, rfl⟩, si⟩, exact ⟨t, rfl, si⟩ }, @@ -1216,12 +1218,12 @@ end /-- A vector space has dimension at most `1` if and only if there is a single vector of which all vectors are multiples. -/ -lemma dim_le_one_iff : module.rank K V ≤ 1 ↔ ∃ v₀ : V, ∀ v, ∃ r : K, r • v₀ = v := +lemma rank_le_one_iff : module.rank K V ≤ 1 ↔ ∃ v₀ : V, ∀ v, ∃ r : K, r • v₀ = v := begin let b := basis.of_vector_space K V, split, { intro hd, - rw [← b.mk_eq_dim'', cardinal.le_one_iff_subsingleton, subsingleton_coe] at hd, + rw [← b.mk_eq_rank'', cardinal.le_one_iff_subsingleton, subsingleton_coe] at hd, rcases eq_empty_or_nonempty (of_vector_space_index K V) with hb | ⟨⟨v₀, hv₀⟩⟩, { use 0, have h' : ∀ v : V, v = 0, { simpa [hb, submodule.eq_bot_iff] using b.span_eq.symm }, @@ -1235,17 +1237,17 @@ begin { rintros ⟨v₀, hv₀⟩, have h : (K ∙ v₀) = ⊤, { ext, simp [mem_span_singleton, hv₀] }, - rw [←dim_top, ←h], - refine (dim_span_le _).trans_eq _, + rw [←rank_top, ←h], + refine (rank_span_le _).trans_eq _, simp } end /-- A submodule has dimension at most `1` if and only if there is a single vector in the submodule such that the submodule is contained in its span. -/ -lemma dim_submodule_le_one_iff (s : submodule K V) : module.rank K s ≤ 1 ↔ ∃ v₀ ∈ s, s ≤ K ∙ v₀ := +lemma rank_submodule_le_one_iff (s : submodule K V) : module.rank K s ≤ 1 ↔ ∃ v₀ ∈ s, s ≤ K ∙ v₀ := begin - simp_rw [dim_le_one_iff, le_span_singleton_iff], + simp_rw [rank_le_one_iff, le_span_singleton_iff], split, { rintro ⟨⟨v₀, hv₀⟩, h⟩, use [v₀, hv₀], @@ -1266,9 +1268,9 @@ end /-- A submodule has dimension at most `1` if and only if there is a single vector, not necessarily in the submodule, such that the submodule is contained in its span. -/ -lemma dim_submodule_le_one_iff' (s : submodule K V) : module.rank K s ≤ 1 ↔ ∃ v₀, s ≤ K ∙ v₀ := +lemma rank_submodule_le_one_iff' (s : submodule K V) : module.rank K s ≤ 1 ↔ ∃ v₀, s ≤ K ∙ v₀ := begin - rw dim_submodule_le_one_iff, + rw rank_submodule_le_one_iff, split, { rintros ⟨v₀, hv₀, h⟩, exact ⟨v₀, h⟩ }, @@ -1289,7 +1291,7 @@ end lemma submodule.rank_le_one_iff_is_principal (W : submodule K V) : module.rank K W ≤ 1 ↔ W.is_principal := begin - simp only [dim_le_one_iff, submodule.is_principal_iff, le_antisymm_iff, + simp only [rank_le_one_iff, submodule.is_principal_iff, le_antisymm_iff, le_span_singleton_iff, span_singleton_le_iff_mem], split, { rintro ⟨⟨m, hm⟩, hm'⟩, @@ -1302,7 +1304,7 @@ end lemma module.rank_le_one_iff_top_is_principal : module.rank K V ≤ 1 ↔ (⊤ : submodule K V).is_principal := -by rw [← submodule.rank_le_one_iff_is_principal, dim_top] +by rw [← submodule.rank_le_one_iff_is_principal, rank_top] end division_ring @@ -1320,16 +1322,16 @@ variables [add_comm_group V'] [module K V'] def rank (f : V →ₗ[K] V') : cardinal := module.rank K f.range lemma rank_le_range (f : V →ₗ[K] V₁) : rank f ≤ module.rank K V₁ := -dim_submodule_le _ +rank_submodule_le _ @[simp] lemma rank_zero [nontrivial K] : rank (0 : V →ₗ[K] V') = 0 := -by rw [rank, linear_map.range_zero, dim_bot] +by rw [rank, linear_map.range_zero, rank_bot] variables [add_comm_group V''] [module K V''] lemma rank_comp_le1 (g : V →ₗ[K] V') (f : V' →ₗ[K] V'') : rank (f.comp g) ≤ rank f := begin - refine dim_le_of_submodule _ _ _, + refine rank_le_of_submodule _ _ _, rw [linear_map.range_comp], exact linear_map.map_le_range, end @@ -1337,7 +1339,7 @@ end variables [add_comm_group V'₁] [module K V'₁] lemma rank_comp_le2 (g : V →ₗ[K] V') (f : V' →ₗ[K] V'₁) : rank (f.comp g) ≤ rank g := -by rw [rank, rank, linear_map.range_comp]; exact dim_map_le _ _ +by rw [rank, rank, linear_map.range_comp]; exact rank_map_le _ _ end ring @@ -1346,17 +1348,17 @@ variables [division_ring K] [add_comm_group V] [module K V] [add_comm_group V₁ variables [add_comm_group V'] [module K V'] lemma rank_le_domain (f : V →ₗ[K] V₁) : rank f ≤ module.rank K V := -by { rw [← dim_range_add_dim_ker f], exact self_le_add_right _ _ } +by { rw [← rank_range_add_rank_ker f], exact self_le_add_right _ _ } lemma rank_add_le (f g : V →ₗ[K] V') : rank (f + g) ≤ rank f + rank g := calc rank (f + g) ≤ module.rank K (f.range ⊔ g.range : submodule K V') : begin - refine dim_le_of_submodule _ _ _, + refine rank_le_of_submodule _ _ _, exact (linear_map.range_le_iff_comap.2 $ eq_top_iff'.2 $ assume x, show f x + g x ∈ (f.range ⊔ g.range : submodule K V'), from mem_sup.2 ⟨_, ⟨x, rfl⟩, _, ⟨x, rfl⟩, rfl⟩) end - ... ≤ rank f + rank g : dim_add_le_dim_add_dim _ _ + ... ≤ rank f + rank g : rank_add_le_rank_add_rank _ _ lemma rank_finset_sum_le {η} (s : finset η) (f : η → V →ₗ[K] V') : rank (∑ d in s, f d) ≤ ∑ d in s, rank (f d) := @@ -1371,7 +1373,7 @@ begin rcases f.range_restrict.exists_right_inverse_of_surjective f.range_range_restrict with ⟨g, hg⟩, have fg : left_inverse f.range_restrict g, from linear_map.congr_fun hg, refine ⟨λ h, _, _⟩, - { rcases le_dim_iff_exists_linear_independent.1 h with ⟨s, rfl, si⟩, + { rcases le_rank_iff_exists_linear_independent.1 h with ⟨s, rfl, si⟩, refine ⟨g '' s, cardinal.mk_image_eq_lift _ _ fg.injective, _⟩, replace fg : ∀ x, f (g x) = x, by { intro x, convert congr_arg subtype.val (fg x) }, replace si : linear_independent K (λ x : s, f (g x)), @@ -1380,7 +1382,7 @@ begin { rintro ⟨s, hsc, si⟩, have : linear_independent K (λ x : s, f.range_restrict x), from linear_independent.of_comp (f.range.subtype) (by convert si), - convert cardinal_le_dim_of_linear_independent this.image, + convert cardinal_le_rank_of_linear_independent this.image, rw [← cardinal.lift_inj, ← hsc, cardinal.mk_image_eq_of_inj_on_lift], exact inj_on_iff_injective.2 this.injective } end diff --git a/src/linear_algebra/dual.lean b/src/linear_algebra/dual.lean index 0244ffc815048..133ca4ffbc02a 100644 --- a/src/linear_algebra/dual.lean +++ b/src/linear_algebra/dual.lean @@ -422,13 +422,13 @@ end comm_ring finsupp.total ι (dual R M) R b.coord f (b i) = f i := by { haveI := classical.dec_eq ι, rw [← coe_dual_basis, total_dual_basis] } -lemma dual_dim_eq [comm_ring K] [add_comm_group V] [module K V] [_root_.finite ι] +lemma dual_rank_eq [comm_ring K] [add_comm_group V] [module K V] [_root_.finite ι] (b : basis ι K V) : cardinal.lift (module.rank K V) = module.rank K (dual K V) := begin classical, casesI nonempty_fintype ι, - have := linear_equiv.lift_dim_eq b.to_dual_equiv, + have := linear_equiv.lift_rank_eq b.to_dual_equiv, simp only [cardinal.lift_umax] at this, rw [this, ← cardinal.lift_umax], apply cardinal.lift_id, @@ -479,9 +479,9 @@ by { rw [← eval_apply_eq_zero_iff K v, linear_map.ext_iff], refl } end -- TODO(jmc): generalize to rings, once `module.rank` is generalized -theorem dual_dim_eq [finite_dimensional K V] : +theorem dual_rank_eq [finite_dimensional K V] : cardinal.lift (module.rank K V) = module.rank K (dual K V) := -(basis.of_vector_space K V).dual_dim_eq +(basis.of_vector_space K V).dual_rank_eq lemma erange_coe [finite_dimensional K V] : (eval K V).range = ⊤ := begin diff --git a/src/linear_algebra/finite_dimensional.lean b/src/linear_algebra/finite_dimensional.lean index 653626d9735df..b2fc8bb9ecb9d 100644 --- a/src/linear_algebra/finite_dimensional.lean +++ b/src/linear_algebra/finite_dimensional.lean @@ -157,7 +157,8 @@ instance finite_dimensional_submodule [finite_dimensional K V] (S : submodule K begin letI : is_noetherian K V := iff_fg.2 _, exact iff_fg.1 - (is_noetherian.iff_dim_lt_aleph_0.2 (lt_of_le_of_lt (dim_submodule_le _) (dim_lt_aleph_0 K V))), + (is_noetherian.iff_rank_lt_aleph_0.2 + (lt_of_le_of_lt (rank_submodule_le _) (rank_lt_aleph_0 K V))), apply_instance, end @@ -169,13 +170,13 @@ module.finite.of_surjective (submodule.mkq S) $ surjective_quot_mk _ variables (K V) /-- In a finite-dimensional space, its dimension (seen as a cardinal) coincides with its `finrank`. This is a copy of `finrank_eq_rank _ _` which creates easier typeclass searches. -/ -lemma finrank_eq_dim [finite_dimensional K V] : +lemma finrank_eq_rank' [finite_dimensional K V] : (finrank K V : cardinal.{v}) = module.rank K V := finrank_eq_rank _ _ variables {K V} lemma finrank_of_infinite_dimensional (h : ¬finite_dimensional K V) : finrank K V = 0 := -dif_neg $ mt is_noetherian.iff_dim_lt_aleph_0.2 $ (not_iff_not.2 iff_fg).2 h +dif_neg $ mt is_noetherian.iff_rank_lt_aleph_0.2 $ (not_iff_not.2 iff_fg).2 h lemma finite_dimensional_of_finrank (h : 0 < finrank K V) : finite_dimensional K V := by { contrapose h, simp [finrank_of_infinite_dimensional h] } @@ -193,7 +194,7 @@ finite_dimensional_of_finrank $ by convert nat.succ_pos n; apply fact.out lemma finite_dimensional_iff_of_rank_eq_nsmul {W} [add_comm_group W] [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_aleph_0, hVW, +by simp only [finite_dimensional, ← is_noetherian.iff_fg, is_noetherian.iff_rank_lt_aleph_0, hVW, cardinal.nsmul_lt_aleph_0_iff_of_ne_zero hn] /-- If a vector space is finite-dimensional, then the cardinality of any basis is equal to its @@ -249,8 +250,8 @@ lemma cardinal_mk_le_finrank_of_linear_independent #ι ≤ finrank K V := begin rw ← lift_le.{_ (max v w)}, - simpa [← finrank_eq_dim, -finrank_eq_rank] using - cardinal_lift_le_dim_of_linear_independent.{_ _ _ (max v w)} h + simpa [← finrank_eq_rank', -finrank_eq_rank] using + cardinal_lift_le_rank_of_linear_independent.{_ _ _ (max v w)} h end lemma fintype_card_le_finrank_of_linear_independent @@ -272,8 +273,8 @@ lemma lt_aleph_0_of_linear_independent {ι : Type w} [finite_dimensional K V] begin apply cardinal.lift_lt.1, apply lt_of_le_of_lt, - apply cardinal_lift_le_dim_of_linear_independent h, - rw [←finrank_eq_dim, cardinal.lift_aleph_0, cardinal.lift_nat_cast], + apply cardinal_lift_le_rank_of_linear_independent h, + rw [←finrank_eq_rank, cardinal.lift_aleph_0, cardinal.lift_nat_cast], apply cardinal.nat_lt_aleph_0, end @@ -292,21 +293,21 @@ end /-- A finite dimensional space has positive `finrank` iff it has a nonzero element. -/ lemma finrank_pos_iff_exists_ne_zero [finite_dimensional K V] : 0 < finrank K V ↔ ∃ x : V, x ≠ 0 := -iff.trans (by { rw ← finrank_eq_dim, norm_cast }) (@dim_pos_iff_exists_ne_zero K V _ _ _ _ _) +iff.trans (by { rw ← finrank_eq_rank, norm_cast }) (@rank_pos_iff_exists_ne_zero K V _ _ _ _ _) /-- A finite dimensional space has positive `finrank` iff it is nontrivial. -/ lemma finrank_pos_iff [finite_dimensional K V] : 0 < finrank K V ↔ nontrivial V := -iff.trans (by { rw ← finrank_eq_dim, norm_cast }) (@dim_pos_iff_nontrivial K V _ _ _ _ _) +iff.trans (by { rw ← finrank_eq_rank, norm_cast }) (@rank_pos_iff_nontrivial K V _ _ _ _ _) /-- A nontrivial finite dimensional space has positive `finrank`. -/ lemma finrank_pos [finite_dimensional K V] [h : nontrivial V] : 0 < finrank K V := finrank_pos_iff.mpr h /-- A finite dimensional space has zero `finrank` iff it is a subsingleton. -This is the `finrank` version of `dim_zero_iff`. -/ +This is the `finrank` version of `rank_zero_iff`. -/ lemma finrank_zero_iff [finite_dimensional K V] : finrank K V = 0 ↔ subsingleton V := -iff.trans (by { rw ← finrank_eq_dim, norm_cast }) (@dim_zero_iff K V _ _ _ _ _) +iff.trans (by { rw ← finrank_eq_rank, norm_cast }) (@rank_zero_iff K V _ _ _ _ _) /-- If a submodule has maximal dimension in a finite dimensional space, then it is equal to the whole space. -/ @@ -358,17 +359,17 @@ span_of_finite K $ s.finite_to_set instance (f : V →ₗ[K] V₂) (p : submodule K V) [h : finite_dimensional K p] : finite_dimensional K (p.map f) := begin - unfreezingI { rw [finite_dimensional, ← iff_fg, is_noetherian.iff_dim_lt_aleph_0] at h ⊢ }, + unfreezingI { rw [finite_dimensional, ← iff_fg, is_noetherian.iff_rank_lt_aleph_0] at h ⊢ }, rw [← cardinal.lift_lt.{v' v}], rw [← cardinal.lift_lt.{v v'}] at h, rw [cardinal.lift_aleph_0] at h ⊢, - exact (lift_dim_map_le f p).trans_lt h + exact (lift_rank_map_le f p).trans_lt h end /-- Pushforwards of finite-dimensional submodules have a smaller finrank. -/ lemma finrank_map_le (f : V →ₗ[K] V₂) (p : submodule K V) [finite_dimensional K p] : finrank K (p.map f) ≤ finrank K p := -by simpa [← finrank_eq_dim, -finrank_eq_rank] using lift_dim_map_le f p +by simpa [← finrank_eq_rank', -finrank_eq_rank] using lift_rank_map_le f p variable {K} @@ -380,7 +381,7 @@ begin { rwa cardinal.lift_le at this }, calc cardinal.lift.{v} (# {i // p i ≠ ⊥}) ≤ cardinal.lift.{w} (module.rank K V) : hp.subtype_ne_bot_le_rank - ... = cardinal.lift.{w} (finrank K V : cardinal.{v}) : by rw finrank_eq_dim + ... = cardinal.lift.{w} (finrank K V : cardinal.{v}) : by rw finrank_eq_rank ... = cardinal.lift.{v} (finrank K V : cardinal.{w}) : by simp end @@ -417,7 +418,7 @@ open finset If a finset has cardinality larger than the dimension of the space, then there is a nontrivial linear relation amongst its elements. -/ -lemma exists_nontrivial_relation_of_dim_lt_card +lemma exists_nontrivial_relation_of_rank_lt_card [finite_dimensional K V] {t : finset V} (h : finrank K V < t.card) : ∃ f : V → K, ∑ e in t, f e • e = 0 ∧ ∃ x ∈ t, f x ≠ 0 := begin @@ -452,7 +453,7 @@ If a finset has cardinality larger than `finrank + 1`, then there is a nontrivial linear relation amongst its elements, such that the coefficients of the relation sum to zero. -/ -lemma exists_nontrivial_relation_sum_zero_of_dim_succ_lt_card +lemma exists_nontrivial_relation_sum_zero_of_rank_succ_lt_card [finite_dimensional K V] {t : finset V} (h : finrank K V + 1 < t.card) : ∃ f : V → K, ∑ e in t, f e • e = 0 ∧ ∑ e in t, f e = 0 ∧ ∃ x ∈ t, f x ≠ 0 := begin @@ -466,7 +467,7 @@ begin { simp only [t', card_map, finset.card_erase_of_mem m], exact nat.lt_pred_iff.mpr h, }, -- to obtain a function `g`. - obtain ⟨g, gsum, x₁, x₁_mem, nz⟩ := exists_nontrivial_relation_of_dim_lt_card h', + obtain ⟨g, gsum, x₁, x₁_mem, nz⟩ := exists_nontrivial_relation_of_rank_lt_card h', -- Then obtain `f` by translating back by `x₀`, -- and setting the value of `f` at `x₀` to ensure `∑ e in t, f e = 0`. let f : V → K := λ z, if z = x₀ then - ∑ z in (t.erase x₀), g (z - x₀) else g (z - x₀), @@ -518,15 +519,15 @@ variables {L : Type*} [linear_ordered_field L] variables {W : Type v} [add_comm_group W] [module L W] /-- -A slight strengthening of `exists_nontrivial_relation_sum_zero_of_dim_succ_lt_card` +A slight strengthening of `exists_nontrivial_relation_sum_zero_of_rank_succ_lt_card` available when working over an ordered field: we can ensure a positive coefficient, not just a nonzero coefficient. -/ -lemma exists_relation_sum_zero_pos_coefficient_of_dim_succ_lt_card +lemma exists_relation_sum_zero_pos_coefficient_of_rank_succ_lt_card [finite_dimensional L W] {t : finset W} (h : finrank L W + 1 < t.card) : ∃ f : W → L, ∑ e in t, f e • e = 0 ∧ ∑ e in t, f e = 0 ∧ ∃ x ∈ t, 0 < f x := begin - obtain ⟨f, sum, total, nonzero⟩ := exists_nontrivial_relation_sum_zero_of_dim_succ_lt_card h, + obtain ⟨f, sum, total, nonzero⟩ := exists_nontrivial_relation_sum_zero_of_rank_succ_lt_card h, exact ⟨f, sum, total, exists_pos_of_sum_zero_of_exists_nonzero f total nonzero⟩, end @@ -576,56 +577,57 @@ end finite_dimensional variables {K V} -section zero_dim +section zero_rank variables [division_ring K] [add_comm_group V] [module K V] open finite_dimensional -lemma finite_dimensional_of_dim_eq_nat {n : ℕ} (h : module.rank K V = n) : finite_dimensional K V := +lemma finite_dimensional_of_rank_eq_nat {n : ℕ} (h : module.rank K V = n) : + finite_dimensional K V := begin - rw [finite_dimensional, ← is_noetherian.iff_fg, is_noetherian.iff_dim_lt_aleph_0, h], + rw [finite_dimensional, ← is_noetherian.iff_fg, is_noetherian.iff_rank_lt_aleph_0, h], exact nat_lt_aleph_0 n, end /- TODO: generalize to free modules over general rings. -/ -lemma finite_dimensional_of_dim_eq_zero (h : module.rank K V = 0) : finite_dimensional K V := -finite_dimensional_of_dim_eq_nat $ h.trans nat.cast_zero.symm +lemma finite_dimensional_of_rank_eq_zero (h : module.rank K V = 0) : finite_dimensional K V := +finite_dimensional_of_rank_eq_nat $ h.trans nat.cast_zero.symm -lemma finite_dimensional_of_dim_eq_one (h : module.rank K V = 1) : finite_dimensional K V := -finite_dimensional_of_dim_eq_nat $ h.trans nat.cast_one.symm +lemma finite_dimensional_of_rank_eq_one (h : module.rank K V = 1) : finite_dimensional K V := +finite_dimensional_of_rank_eq_nat $ h.trans nat.cast_one.symm -lemma finrank_eq_zero_of_dim_eq_zero [finite_dimensional K V] (h : module.rank K V = 0) : +lemma finrank_eq_zero_of_rank_eq_zero [finite_dimensional K V] (h : module.rank K V = 0) : finrank K V = 0 := begin - convert finrank_eq_dim K V, + convert finrank_eq_rank K V, rw h, norm_cast end variables (K V) instance finite_dimensional_bot : finite_dimensional K (⊥ : submodule K V) := -finite_dimensional_of_dim_eq_zero $ by simp +finite_dimensional_of_rank_eq_zero $ by simp variables {K V} -lemma bot_eq_top_of_dim_eq_zero (h : module.rank K V = 0) : (⊥ : submodule K V) = ⊤ := +lemma bot_eq_top_of_rank_eq_zero (h : module.rank K V = 0) : (⊥ : submodule K V) = ⊤ := begin - haveI := finite_dimensional_of_dim_eq_zero h, + haveI := finite_dimensional_of_rank_eq_zero h, apply eq_top_of_finrank_eq, - rw [finrank_bot, finrank_eq_zero_of_dim_eq_zero h] + rw [finrank_bot, finrank_eq_zero_of_rank_eq_zero h] end -@[simp] theorem dim_eq_zero {S : submodule K V} : module.rank K S = 0 ↔ S = ⊥ := +@[simp] theorem rank_eq_zero {S : submodule K V} : module.rank K S = 0 ↔ S = ⊥ := ⟨λ h, (submodule.eq_bot_iff _).2 $ λ x hx, congr_arg subtype.val $ - ((submodule.eq_bot_iff _).1 $ eq.symm $ bot_eq_top_of_dim_eq_zero h) ⟨x, hx⟩ submodule.mem_top, -λ h, by rw [h, dim_bot]⟩ + ((submodule.eq_bot_iff _).1 $ eq.symm $ bot_eq_top_of_rank_eq_zero h) ⟨x, hx⟩ submodule.mem_top, +λ h, by rw [h, rank_bot]⟩ @[simp] theorem finrank_eq_zero {S : submodule K V} [finite_dimensional K S] : finrank K S = 0 ↔ S = ⊥ := -by rw [← dim_eq_zero, ← finrank_eq_dim, ← @nat.cast_zero cardinal, cardinal.nat_cast_inj] +by rw [← rank_eq_zero, ← finrank_eq_rank, ← @nat.cast_zero cardinal, cardinal.nat_cast_inj] -end zero_dim +end zero_rank namespace submodule open is_noetherian finite_dimensional @@ -644,8 +646,8 @@ lemma finite_dimensional_of_le {S₁ S₂ : submodule K V} [finite_dimensional K finite_dimensional K S₁ := begin haveI : is_noetherian K S₂ := iff_fg.2 infer_instance, - exact iff_fg.1 (is_noetherian.iff_dim_lt_aleph_0.2 - (lt_of_le_of_lt (dim_le_of_submodule _ _ h) (dim_lt_aleph_0 K S₂))), + exact iff_fg.1 (is_noetherian.iff_rank_lt_aleph_0.2 + (lt_of_le_of_lt (rank_le_of_submodule _ _ h) (finite_dimensional.rank_lt_aleph_0 K S₂))), end /-- The inf of two submodules, the first finite-dimensional, is @@ -696,22 +698,22 @@ 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) +by simpa only [cardinal.nat_cast_le, ←finrank_eq_rank] using + s.subtype.rank_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 _) +by simpa only [cardinal.nat_cast_le, ←finrank_eq_rank] using + (mkq s).rank_le_of_surjective (surjective_quot_mk _) /-- 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) : finrank K (V ⧸ s) + finrank K s = finrank K V := begin - have := dim_quotient_add_dim s, - rw [← finrank_eq_dim, ← finrank_eq_dim, ← finrank_eq_dim] at this, + have := rank_quotient_add_rank s, + rw [← finrank_eq_rank, ← finrank_eq_rank, ← finrank_eq_rank] at this, exact_mod_cast this end @@ -725,21 +727,21 @@ begin end /-- 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) +theorem rank_sup_add_rank_inf_eq (s t : submodule K V) [finite_dimensional K s] [finite_dimensional K t] : finrank K ↥(s ⊔ t) + finrank K ↥(s ⊓ t) = finrank K ↥s + finrank K ↥t := begin have key : module.rank K ↥(s ⊔ t) + module.rank K ↥(s ⊓ t) = - module.rank K s + module.rank K t := dim_sup_add_dim_inf_eq s t, - repeat { rw ←finrank_eq_dim at key }, + module.rank K s + module.rank K t := rank_sup_add_rank_inf_eq s t, + repeat { rw ←finrank_eq_rank at key }, norm_cast at key, exact key end -lemma dim_add_le_dim_add_dim (s t : submodule K V) +lemma rank_add_le_rank_add_rank (s t : submodule K V) [finite_dimensional K s] [finite_dimensional K t] : finrank K (s ⊔ t : submodule K V) ≤ finrank K s + finrank K t := -by { rw [← dim_sup_add_dim_inf_eq], exact self_le_add_right _ _ } +by { rw [← rank_sup_add_rank_inf_eq], exact self_le_add_right _ _ } lemma eq_top_of_disjoint [finite_dimensional K V] (s t : submodule K V) (hdim : finrank K s + finrank K t = finrank K V) @@ -750,7 +752,7 @@ begin rw [hdisjoint, finrank_bot] }, apply eq_top_of_finrank_eq, rw ←hdim, - convert s.dim_sup_add_dim_inf_eq t, + convert s.rank_sup_add_rank_inf_eq t, rw h_finrank_inf, refl, end @@ -795,7 +797,7 @@ Two finite-dimensional vector spaces are isomorphic if they have the same (finit -/ theorem nonempty_linear_equiv_of_finrank_eq [finite_dimensional K V] [finite_dimensional K V₂] (cond : finrank K V = finrank K V₂) : nonempty (V ≃ₗ[K] V₂) := -nonempty_linear_equiv_of_lift_dim_eq $ by simp only [← finrank_eq_dim, cond, lift_nat_cast] +nonempty_linear_equiv_of_lift_rank_eq $ by simp only [← finrank_eq_rank, cond, lift_nat_cast] /-- Two finite-dimensional vector spaces are isomorphic if and only if they have the same (finite) @@ -871,8 +873,8 @@ variables [division_ring K] [add_comm_group V] [module K V] lemma surjective_of_injective [finite_dimensional K V] {f : V →ₗ[K] V} (hinj : injective f) : surjective f := begin - have h := dim_eq_of_injective _ hinj, - rw [← finrank_eq_dim, ← finrank_eq_dim, nat_cast_inj] at h, + have h := rank_eq_of_injective _ hinj, + rw [← finrank_eq_rank, ← finrank_eq_rank, nat_cast_inj] at h, exact range_eq_top.1 (eq_top_of_finrank_eq h.symm) end @@ -1092,7 +1094,7 @@ lemma finrank_add_eq_of_is_compl [finite_dimensional K V] {U W : submodule K V} (h : is_compl U W) : finrank K U + finrank K W = finrank K V := begin - rw [← dim_sup_add_dim_inf_eq, h.codisjoint.eq_top, h.disjoint.eq_bot, finrank_bot, add_zero], + rw [← rank_sup_add_rank_inf_eq, h.codisjoint.eq_top, h.disjoint.eq_bot, finrank_bot, add_zero], exact finrank_top end @@ -1138,7 +1140,7 @@ begin symmetry, replace fin := (not_iff_not.2 is_noetherian.iff_fg).2 fin, calc fintype.card ι = finrank K V : card_eq - ... = 0 : dif_neg (mt is_noetherian.iff_dim_lt_aleph_0.mpr fin) } + ... = 0 : dif_neg (mt is_noetherian.iff_rank_lt_aleph_0.mpr fin) } end /-- A linear independent family of `finrank K V` vectors forms a basis. -/ @@ -1260,11 +1262,11 @@ end lemma submodule.finrank_le_one_iff_is_principal (W : submodule K V) [finite_dimensional K W] : finrank K W ≤ 1 ↔ W.is_principal := -by rw [← W.rank_le_one_iff_is_principal, ← finrank_eq_dim, ← cardinal.nat_cast_le, nat.cast_one] +by rw [← W.rank_le_one_iff_is_principal, ← finrank_eq_rank, ← cardinal.nat_cast_le, nat.cast_one] lemma module.finrank_le_one_iff_top_is_principal [finite_dimensional K V] : finrank K V ≤ 1 ↔ (⊤ : submodule K V).is_principal := -by rw [← module.rank_le_one_iff_top_is_principal, ← finrank_eq_dim, +by rw [← module.rank_le_one_iff_top_is_principal, ← finrank_eq_rank, ← cardinal.nat_cast_le, nat.cast_one] -- We use the `linear_map.compatible_smul` typeclass here, to encompass two situations: @@ -1297,7 +1299,7 @@ end finrank_eq_one end division_ring -section subalgebra_dim +section subalgebra_rank open module variables {F E : Type*} [field F] [ring E] [algebra F E] @@ -1313,46 +1315,46 @@ instance finite_dimensional.finite_dimensional_subalgebra [finite_dimensional F finite_dimensional.of_subalgebra_to_submodule infer_instance instance subalgebra.finite_dimensional_bot : finite_dimensional F (⊥ : subalgebra F E) := -by { nontriviality E, exact finite_dimensional_of_dim_eq_one subalgebra.dim_bot } +by { nontriviality E, exact finite_dimensional_of_rank_eq_one subalgebra.rank_bot } -lemma subalgebra.eq_bot_of_dim_le_one {S : subalgebra F E} (h : module.rank F S ≤ 1) : S = ⊥ := +lemma subalgebra.eq_bot_of_rank_le_one {S : subalgebra F E} (h : module.rank F S ≤ 1) : S = ⊥ := begin nontriviality E, obtain ⟨m, hm, he⟩ := cardinal.exists_nat_eq_of_le_nat (h.trans_eq nat.cast_one.symm), - haveI := finite_dimensional_of_dim_eq_nat he, + haveI := finite_dimensional_of_rank_eq_nat he, rw [← not_bot_lt_iff, ← subalgebra.to_submodule.lt_iff_lt], haveI := (S.to_submodule_equiv).symm.finite_dimensional, refine λ hl, (submodule.finrank_lt_finrank_of_lt hl).not_le (nat_cast_le.1 _), - iterate 2 { rw [subalgebra.finrank_to_submodule, finrank_eq_dim] }, - exact h.trans_eq subalgebra.dim_bot.symm, + iterate 2 { rw [subalgebra.finrank_to_submodule, finrank_eq_rank] }, + exact h.trans_eq subalgebra.rank_bot.symm, end lemma subalgebra.eq_bot_of_finrank_one {S : subalgebra F E} (h : finrank F S = 1) : S = ⊥ := -subalgebra.eq_bot_of_dim_le_one $ - by { haveI := finite_dimensional_of_finrank_eq_succ h, rw [← finrank_eq_dim, h, nat.cast_one] } +subalgebra.eq_bot_of_rank_le_one $ + by { haveI := finite_dimensional_of_finrank_eq_succ h, rw [← finrank_eq_rank, h, nat.cast_one] } @[simp] -theorem subalgebra.dim_eq_one_iff [nontrivial E] {S : subalgebra F E} : +theorem subalgebra.rank_eq_one_iff [nontrivial E] {S : subalgebra F E} : module.rank F S = 1 ↔ S = ⊥ := -⟨λ h, subalgebra.eq_bot_of_dim_le_one h.le, λ h, h.symm ▸ subalgebra.dim_bot⟩ +⟨λ h, subalgebra.eq_bot_of_rank_le_one h.le, λ h, h.symm ▸ subalgebra.rank_bot⟩ @[simp] theorem subalgebra.finrank_eq_one_iff [nontrivial E] {S : subalgebra F E} : finrank F S = 1 ↔ S = ⊥ := ⟨subalgebra.eq_bot_of_finrank_one, λ h, h.symm ▸ subalgebra.finrank_bot⟩ -lemma subalgebra.bot_eq_top_iff_dim_eq_one [nontrivial E] : +lemma subalgebra.bot_eq_top_iff_rank_eq_one [nontrivial E] : (⊥ : subalgebra F E) = ⊤ ↔ module.rank F E = 1 := -by rw [← dim_top, ← subalgebra_top_dim_eq_submodule_top_dim, subalgebra.dim_eq_one_iff, eq_comm] +by rw [← rank_top, ← subalgebra_top_rank_eq_submodule_top_rank, subalgebra.rank_eq_one_iff, eq_comm] lemma subalgebra.bot_eq_top_iff_finrank_eq_one [nontrivial E] : (⊥ : subalgebra F E) = ⊤ ↔ finrank F E = 1 := by rw [← finrank_top, ← subalgebra_top_finrank_eq_submodule_top_finrank, subalgebra.finrank_eq_one_iff, eq_comm] -alias subalgebra.bot_eq_top_iff_dim_eq_one ↔ _ subalgebra.bot_eq_top_of_dim_eq_one +alias subalgebra.bot_eq_top_iff_rank_eq_one ↔ _ subalgebra.bot_eq_top_of_rank_eq_one alias subalgebra.bot_eq_top_iff_finrank_eq_one ↔ _ subalgebra.bot_eq_top_of_finrank_eq_one -attribute [simp] subalgebra.bot_eq_top_of_finrank_eq_one subalgebra.bot_eq_top_of_dim_eq_one +attribute [simp] subalgebra.bot_eq_top_of_finrank_eq_one subalgebra.bot_eq_top_of_rank_eq_one lemma subalgebra.is_simple_order_of_finrank (hr : finrank F E = 2) : is_simple_order (subalgebra F E) := @@ -1374,7 +1376,7 @@ let i := nontrivial_of_finrank_pos (zero_lt_two.trans_eq hr.symm) in by exactI exact submodule.eq_top_of_finrank_eq h, }, end } -end subalgebra_dim +end subalgebra_rank namespace module namespace End @@ -1457,7 +1459,7 @@ open module open_locale cardinal -lemma cardinal_mk_eq_cardinal_mk_field_pow_dim +lemma cardinal_mk_eq_cardinal_mk_field_pow_rank (K V : Type u) [division_ring K] [add_comm_group V] [module K V] [finite_dimensional K V] : #V = #K ^ module.rank K V := begin @@ -1465,7 +1467,7 @@ begin let hs := basis.of_vector_space K V, calc #V = #(s →₀ K) : quotient.sound ⟨hs.repr.to_equiv⟩ ... = #(s → K) : quotient.sound ⟨finsupp.equiv_fun_on_finite⟩ - ... = _ : by rw [← cardinal.lift_inj.1 hs.mk_eq_dim, cardinal.power_def] + ... = _ : by rw [← cardinal.lift_inj.1 hs.mk_eq_rank, cardinal.power_def] end lemma cardinal_lt_aleph_0_of_finite_dimensional @@ -1474,9 +1476,9 @@ lemma cardinal_lt_aleph_0_of_finite_dimensional #V < ℵ₀ := begin letI : is_noetherian K V := is_noetherian.iff_fg.2 infer_instance, - rw cardinal_mk_eq_cardinal_mk_field_pow_dim K V, + rw cardinal_mk_eq_cardinal_mk_field_pow_rank K V, exact cardinal.power_lt_aleph_0 (cardinal.lt_aleph_0_of_finite K) - (is_noetherian.dim_lt_aleph_0 K V), + (is_noetherian.rank_lt_aleph_0 K V), end end module diff --git a/src/linear_algebra/finrank.lean b/src/linear_algebra/finrank.lean index e11f3a76d96f5..5aaafab257c8c 100644 --- a/src/linear_algebra/finrank.lean +++ b/src/linear_algebra/finrank.lean @@ -56,28 +56,28 @@ noncomputable def finrank (R V : Type*) [semiring R] [add_comm_group V] [module R V] : ℕ := (module.rank R V).to_nat -lemma finrank_eq_of_dim_eq {n : ℕ} (h : module.rank K V = ↑ n) : finrank K V = n := +lemma finrank_eq_of_rank_eq {n : ℕ} (h : module.rank K V = ↑ n) : finrank K V = n := begin apply_fun to_nat at h, rw to_nat_cast at h, exact_mod_cast h, end -lemma finrank_le_of_dim_le {n : ℕ} (h : module.rank K V ≤ ↑ n) : finrank K V ≤ n := +lemma finrank_le_of_rank_le {n : ℕ} (h : module.rank K V ≤ ↑ n) : finrank K V ≤ n := begin rwa [← cardinal.to_nat_le_iff_le_of_lt_aleph_0, to_nat_cast] at h, { exact h.trans_lt (nat_lt_aleph_0 n) }, { exact nat_lt_aleph_0 n }, end -lemma finrank_lt_of_dim_lt {n : ℕ} (h : module.rank K V < ↑ n) : finrank K V < n := +lemma finrank_lt_of_rank_lt {n : ℕ} (h : module.rank K V < ↑ n) : finrank K V < n := begin rwa [← cardinal.to_nat_lt_iff_lt_of_lt_aleph_0, to_nat_cast] at h, { exact h.trans (nat_lt_aleph_0 n) }, { exact nat_lt_aleph_0 n }, end -lemma dim_lt_of_finrank_lt {n : ℕ} (h : n < finrank K V) : ↑n < module.rank K V := +lemma rank_lt_of_finrank_lt {n : ℕ} (h : n < finrank K V) : ↑n < module.rank K V := begin rwa [← cardinal.to_nat_lt_iff_lt_of_lt_aleph_0, to_nat_cast], { exact nat_lt_aleph_0 n }, @@ -90,7 +90,7 @@ end basis. -/ lemma finrank_eq_card_basis {ι : Type w} [fintype ι] (h : basis ι K V) : finrank K V = fintype.card ι := -finrank_eq_of_dim_eq (dim_eq_card_basis h) +finrank_eq_of_rank_eq (rank_eq_card_basis h) /-- If a vector space has a finite basis, then its dimension is equal to the cardinality of the basis. This lemma uses a `finset` instead of indexed types. -/ @@ -101,7 +101,7 @@ by rw [finrank_eq_card_basis h, fintype.card_coe] /-- A finite dimensional space is nontrivial if it has positive `finrank`. -/ lemma nontrivial_of_finrank_pos (h : 0 < finrank K V) : nontrivial V := -dim_pos_iff_nontrivial.mp (dim_lt_of_finrank_lt h) +rank_pos_iff_nontrivial.mp (rank_lt_of_finrank_lt h) /-- A finite dimensional space is nontrivial if it has `finrank` equal to the successor of a natural number. -/ @@ -124,12 +124,12 @@ hs.subset_extend _ variable (K) /-- A division_ring is one-dimensional as a vector space over itself. -/ @[simp] lemma finrank_self : finrank K K = 1 := -finrank_eq_of_dim_eq (by simp) +finrank_eq_of_rank_eq (by simp) /-- The vector space of functions on a fintype ι has finrank equal to the cardinality of ι. -/ @[simp] lemma finrank_fintype_fun_eq_card {ι : Type v} [fintype ι] : finrank K (ι → K) = fintype.card ι := -finrank_eq_of_dim_eq dim_fun' +finrank_eq_of_rank_eq rank_fun' /-- The vector space of functions on `fin n` has finrank equal to `n`. -/ @[simp] lemma finrank_fin_fun {n : ℕ} : finrank K (fin n → K) = n := @@ -141,7 +141,7 @@ end finite_dimensional variables {K V} -section zero_dim +section zero_rank variables [division_ring K] [add_comm_group V] [module K V] @@ -149,8 +149,8 @@ open finite_dimensional lemma finrank_eq_zero_of_basis_imp_not_finite (h : ∀ s : set V, basis.{v} (s : set V) K V → ¬ s.finite) : finrank K V = 0 := -dif_neg (λ dim_lt, h _ (basis.of_vector_space K V) - ((basis.of_vector_space K V).finite_index_of_dim_lt_aleph_0 dim_lt)) +dif_neg (λ rank_lt, h _ (basis.of_vector_space K V) + ((basis.of_vector_space K V).finite_index_of_rank_lt_aleph_0 rank_lt)) lemma finrank_eq_zero_of_basis_imp_false (h : ∀ s : finset V, basis.{v} (s : set V) K V → false) : finrank K V = 0 := @@ -171,9 +171,9 @@ finrank_eq_zero_of_basis_imp_false (λ s b, h ⟨s, ⟨b⟩⟩) variables (K V) @[simp] lemma finrank_bot : finrank K (⊥ : submodule K V) = 0 := -finrank_eq_of_dim_eq (dim_bot _ _) +finrank_eq_of_rank_eq (rank_bot _ _) -end zero_dim +end zero_rank namespace linear_equiv open finite_dimensional @@ -186,7 +186,7 @@ variables [module R M] [module R M₂] /-- The dimension of a finite dimensional space is preserved under linear equivalence. -/ theorem finrank_eq (f : M ≃ₗ[R] M₂) : finrank R M = finrank R M₂ := -by { unfold finrank, rw [← cardinal.to_nat_lift, f.lift_dim_eq, cardinal.to_nat_lift] } +by { unfold finrank, rw [← cardinal.to_nat_lift, f.lift_rank_eq, cardinal.to_nat_lift] } /-- Pushforwards of finite-dimensional submodules along a `linear_equiv` have the same finrank. -/ lemma finrank_map_eq (f : M ≃ₗ[R] M₂) (p : submodule R M) : @@ -218,7 +218,7 @@ variables [division_ring K] [add_comm_group V] [module K V] @[simp] theorem finrank_top : finrank K (⊤ : submodule K V) = finrank K V := -by { unfold finrank, simp [dim_top] } +by { unfold finrank, simp [rank_top] } end @@ -259,7 +259,7 @@ variable {K} lemma finrank_span_le_card (s : set V) [fintype s] : finrank K (span K s) ≤ s.to_finset.card := -finrank_le_of_dim_le (by simpa using dim_span_le s) +finrank_le_of_rank_le (by simpa using rank_span_le s) lemma finrank_span_finset_le_card (s : finset V) : (s : set V).finrank K ≤ s.card := @@ -273,9 +273,9 @@ lemma finrank_range_le_card {ι : Type*} [fintype ι] {b : ι → V} : lemma finrank_span_eq_card {ι : Type*} [fintype ι] {b : ι → V} (hb : linear_independent K b) : finrank K (span K (set.range b)) = fintype.card ι := -finrank_eq_of_dim_eq +finrank_eq_of_rank_eq begin - have : module.rank K (span K (set.range b)) = #(set.range b) := dim_span hb, + have : module.rank K (span K (set.range b)) = #(set.range b) := rank_span hb, rwa [←lift_inj, mk_range_eq_of_injective hb.injective, cardinal.mk_fintype, lift_nat_cast, lift_eq_nat_iff] at this, end @@ -283,9 +283,9 @@ end lemma finrank_span_set_eq_card (s : set V) [fintype s] (hs : linear_independent K (coe : s → V)) : finrank K (span K s) = s.to_finset.card := -finrank_eq_of_dim_eq +finrank_eq_of_rank_eq begin - have : module.rank K (span K s) = #s := dim_span_set hs, + have : module.rank K (span K s) = #s := rank_span_set hs, rwa [cardinal.mk_fintype, ←set.to_finset_card] at this, end @@ -456,22 +456,22 @@ end end finrank_eq_one -section subalgebra_dim +section subalgebra_rank open module variables {F E : Type*} [field F] [ring E] [algebra F E] -@[simp] lemma subalgebra.dim_bot [nontrivial E] : module.rank F (⊥ : subalgebra F E) = 1 := +@[simp] lemma subalgebra.rank_bot [nontrivial E] : module.rank F (⊥ : subalgebra F E) = 1 := ((subalgebra.to_submodule_equiv (⊥ : subalgebra F E)).symm.trans $ - linear_equiv.of_eq _ _ algebra.to_submodule_bot).dim_eq.trans $ - by { rw dim_span_set, exacts [mk_singleton _, linear_independent_singleton one_ne_zero] } + linear_equiv.of_eq _ _ algebra.to_submodule_bot).rank_eq.trans $ + by { rw rank_span_set, exacts [mk_singleton _, linear_independent_singleton one_ne_zero] } -@[simp] lemma subalgebra.dim_to_submodule (S : subalgebra F E) : +@[simp] lemma subalgebra.rank_to_submodule (S : subalgebra F E) : module.rank F S.to_submodule = module.rank F S := rfl @[simp] lemma subalgebra.finrank_to_submodule (S : subalgebra F E) : finrank F S.to_submodule = finrank F S := rfl -lemma subalgebra_top_dim_eq_submodule_top_dim : +lemma subalgebra_top_rank_eq_submodule_top_rank : module.rank F (⊤ : subalgebra F E) = module.rank F (⊤ : submodule F E) := by { rw ← algebra.top_to_submodule, refl } @@ -479,11 +479,11 @@ lemma subalgebra_top_finrank_eq_submodule_top_finrank : finrank F (⊤ : subalgebra F E) = finrank F (⊤ : submodule F E) := by { rw ← algebra.top_to_submodule, refl } -lemma subalgebra.dim_top : module.rank F (⊤ : subalgebra F E) = module.rank F E := -by { rw subalgebra_top_dim_eq_submodule_top_dim, exact dim_top F E } +lemma subalgebra.rank_top : module.rank F (⊤ : subalgebra F E) = module.rank F E := +by { rw subalgebra_top_rank_eq_submodule_top_rank, exact rank_top F E } @[simp] lemma subalgebra.finrank_bot [nontrivial E] : finrank F (⊥ : subalgebra F E) = 1 := -finrank_eq_of_dim_eq (by simp) +finrank_eq_of_rank_eq (by simp) -end subalgebra_dim +end subalgebra_rank diff --git a/src/linear_algebra/free_algebra.lean b/src/linear_algebra/free_algebra.lean index 4e805ddeda500..409c4c0c8154a 100644 --- a/src/linear_algebra/free_algebra.lean +++ b/src/linear_algebra/free_algebra.lean @@ -27,8 +27,8 @@ finsupp.basis_single_one.map (equiv_monoid_algebra_free_monoid.symm.to_linear_equiv : _ ≃ₗ[R] free_algebra R X) -- TODO: generalize to `X : Type v` -lemma dim_eq {K : Type u} {X : Type (max u v)} [field K] : +lemma rank_eq {K : Type u} {X : Type (max u v)} [field K] : module.rank K (free_algebra K X) = cardinal.mk (list X) := -(cardinal.lift_inj.mp (basis_free_monoid K X).mk_eq_dim).symm +(cardinal.lift_inj.mp (basis_free_monoid K X).mk_eq_rank).symm end free_algebra diff --git a/src/linear_algebra/free_module/finite/matrix.lean b/src/linear_algebra/free_module/finite/matrix.lean index 90590c37e3353..4883618aa831b 100644 --- a/src/linear_algebra/free_module/finite/matrix.lean +++ b/src/linear_algebra/free_module/finite/matrix.lean @@ -87,7 +87,7 @@ begin letI := nontrivial_of_invariant_basis_number R, have h := (linear_map.to_matrix (choose_basis R M) (choose_basis R N)), let b := (matrix.std_basis _ _ _).map h.symm, - rw [finrank, dim_eq_card_basis b, ← cardinal.mk_fintype, cardinal.mk_to_nat_eq_card, finrank, + rw [finrank, rank_eq_card_basis b, ← cardinal.mk_fintype, cardinal.mk_to_nat_eq_card, finrank, finrank, rank_eq_card_choose_basis_index, rank_eq_card_choose_basis_index, cardinal.mk_to_nat_eq_card, cardinal.mk_to_nat_eq_card, fintype.card_prod, mul_comm] end diff --git a/src/linear_algebra/free_module/finite/rank.lean b/src/linear_algebra/free_module/finite/rank.lean index 8edb37749b494..9e801be89d6fd 100644 --- a/src/linear_algebra/free_module/finite/rank.lean +++ b/src/linear_algebra/free_module/finite/rank.lean @@ -39,7 +39,7 @@ variables [add_comm_group N] [module R N] [module.free R N] [module.finite R N] lemma rank_lt_aleph_0 : module.rank R M < ℵ₀ := begin letI := nontrivial_of_invariant_basis_number R, - rw [← (choose_basis R M).mk_eq_dim'', lt_aleph_0_iff_fintype], + rw [← (choose_basis R M).mk_eq_rank'', lt_aleph_0_iff_fintype], exact nonempty.intro infer_instance end diff --git a/src/linear_algebra/free_module/rank.lean b/src/linear_algebra/free_module/rank.lean index 6a7379b9c11d2..5eb9e26f3617c 100644 --- a/src/linear_algebra/free_module/rank.lean +++ b/src/linear_algebra/free_module/rank.lean @@ -34,12 +34,12 @@ variables [add_comm_group N] [module R N] [module.free R N] /-- The rank of a free module `M` over `R` is the cardinality of `choose_basis_index R M`. -/ lemma rank_eq_card_choose_basis_index : module.rank R M = #(choose_basis_index R M) := -(choose_basis R M).mk_eq_dim''.symm +(choose_basis R M).mk_eq_rank''.symm /-- The rank of `(ι →₀ R)` is `(# ι).lift`. -/ @[simp] lemma rank_finsupp {ι : Type v} : module.rank R (ι →₀ R) = (# ι).lift := by simpa [lift_id', lift_umax] using - (basis.of_repr (linear_equiv.refl _ (ι →₀ R))).mk_eq_dim.symm + (basis.of_repr (linear_equiv.refl _ (ι →₀ R))).mk_eq_rank.symm /-- If `R` and `ι` lie in the same universe, the rank of `(ι →₀ R)` is `# ι`. -/ lemma rank_finsupp' {ι : Type u} : module.rank R (ι →₀ R) = # ι := by simp @@ -48,7 +48,7 @@ lemma rank_finsupp' {ι : Type u} : module.rank R (ι →₀ R) = # ι := by sim @[simp] lemma rank_prod : module.rank R (M × N) = lift.{w v} (module.rank R M) + lift.{v w} (module.rank R N) := by simpa [rank_eq_card_choose_basis_index R M, rank_eq_card_choose_basis_index R N, - lift_umax, lift_umax'] using ((choose_basis R M).prod (choose_basis R N)).mk_eq_dim.symm + lift_umax, lift_umax'] using ((choose_basis R M).prod (choose_basis R N)).mk_eq_rank.symm /-- If `M` and `N` lie in the same universe, the rank of `M × N` is `(module.rank R M) + (module.rank R N)`. -/ @@ -62,7 +62,7 @@ lemma rank_prod' (N : Type v) [add_comm_group N] [module R N] [module.free R N] begin let B := λ i, choose_basis R (M i), let b : basis _ R (⨁ i, M i) := dfinsupp.basis (λ i, B i), - simp [← b.mk_eq_dim'', λ i, (B i).mk_eq_dim''], + simp [← b.mk_eq_rank'', λ i, (B i).mk_eq_rank''], end /-- The rank of a finite product is the sum of the ranks. -/ @@ -70,7 +70,7 @@ end [Π (i : ι), add_comm_group (M i)] [Π (i : ι), module R (M i)] [Π (i : ι), module.free R (M i)] : module.rank R (Π i, M i) = cardinal.sum (λ i, module.rank R (M i)) := by { casesI nonempty_fintype ι, - rw [←(direct_sum.linear_equiv_fun_on_fintype _ _ M).dim_eq, rank_direct_sum] } + rw [←(direct_sum.linear_equiv_fun_on_fintype _ _ M).rank_eq, rank_direct_sum] } /-- If `m` and `n` are `fintype`, the rank of `m × n` matrices is `(# m).lift * (# n).lift`. -/ @[simp] lemma rank_matrix (m : Type v) (n : Type w) [finite m] [finite n] : @@ -78,7 +78,7 @@ by { casesI nonempty_fintype ι, begin casesI nonempty_fintype m, casesI nonempty_fintype n, - have h := (matrix.std_basis R m n).mk_eq_dim, + have h := (matrix.std_basis R m n).mk_eq_rank, rw [← lift_lift.{(max v w u) (max v w)}, lift_inj] at h, simpa using h.symm, end @@ -109,9 +109,9 @@ begin let ιM := choose_basis_index R M, let ιN := choose_basis_index R N, - have h₁ := linear_equiv.lift_dim_eq (tensor_product.congr (repr R M) (repr R N)), + have h₁ := linear_equiv.lift_rank_eq (tensor_product.congr (repr R M) (repr R N)), let b : basis (ιM × ιN) R (_ →₀ R) := finsupp.basis_single_one, - rw [linear_equiv.dim_eq (finsupp_tensor_finsupp' R ιM ιN), ← b.mk_eq_dim, mk_prod] at h₁, + rw [linear_equiv.rank_eq (finsupp_tensor_finsupp' R ιM ιN), ← b.mk_eq_rank, mk_prod] at h₁, rw [lift_inj.1 h₁, rank_eq_card_choose_basis_index R M, rank_eq_card_choose_basis_index R N], end diff --git a/src/linear_algebra/matrix/diagonal.lean b/src/linear_algebra/matrix/diagonal.lean index 74d5e149b0176..b46dce0d98a54 100644 --- a/src/linear_algebra/matrix/diagonal.lean +++ b/src/linear_algebra/matrix/diagonal.lean @@ -77,8 +77,8 @@ begin have hd : disjoint {i : m | w i ≠ 0} {i : m | w i = 0} := disjoint_compl_left, have B₁ := supr_range_std_basis_eq_infi_ker_proj K (λi:m, K) hd hu (set.to_finite _), have B₂ := @infi_ker_proj_equiv K _ _ (λi:m, K) _ _ _ _ (by simp; apply_instance) hd hu, - rw [rank, range_diagonal, B₁, ←@dim_fun' K], - apply linear_equiv.dim_eq, + rw [rank, range_diagonal, B₁, ←@rank_fun' K], + apply linear_equiv.rank_eq, apply B₂, end diff --git a/src/linear_algebra/matrix/to_lin.lean b/src/linear_algebra/matrix/to_lin.lean index d9b4d10c6e987..380fa4e5b4e41 100644 --- a/src/linear_algebra/matrix/to_lin.lean +++ b/src/linear_algebra/matrix/to_lin.lean @@ -349,7 +349,7 @@ begin rw [vec_mul_vec_eq, matrix.to_lin'_mul], refine le_trans (rank_comp_le1 _ _) _, refine (rank_le_domain _).trans_eq _, - rw [dim_fun', fintype.card_unit, nat.cast_one] + rw [rank_fun', fintype.card_unit, nat.cast_one] end end to_matrix' diff --git a/src/number_theory/ramification_inertia.lean b/src/number_theory/ramification_inertia.lean index 2e66725f662bc..8c61439b6981f 100644 --- a/src/number_theory/ramification_inertia.lean +++ b/src/number_theory/ramification_inertia.lean @@ -606,44 +606,44 @@ end /-- Since the inclusion `(P^(i + 1) / P^e) ⊂ (P^i / P^e)` has a kernel isomorphic to `P / S`, `[P^i / P^e : R / p] = [P^(i+1) / P^e : R / p] + [P / S : R / p]` -/ -lemma dim_pow_quot_aux [is_domain S] [is_dedekind_domain S] [p.is_maximal] [P.is_prime] +lemma rank_pow_quot_aux [is_domain S] [is_dedekind_domain S] [p.is_maximal] [P.is_prime] (hP0 : P ≠ ⊥) {i : ℕ} (hi : i < e) : module.rank (R ⧸ p) (ideal.map (P^e)^.quotient.mk (P ^ i)) = module.rank (R ⧸ p) (S ⧸ P) + module.rank (R ⧸ p) (ideal.map (P^e)^.quotient.mk (P ^ (i + 1))) := begin letI : field (R ⧸ p) := ideal.quotient.field _, - rw [dim_eq_of_injective _ (pow_quot_succ_inclusion_injective f p P i), - (quotient_range_pow_quot_succ_inclusion_equiv f p P hP0 hi).symm.dim_eq], - exact (dim_quotient_add_dim (linear_map.range (pow_quot_succ_inclusion f p P i))).symm, + rw [rank_eq_of_injective _ (pow_quot_succ_inclusion_injective f p P i), + (quotient_range_pow_quot_succ_inclusion_equiv f p P hP0 hi).symm.rank_eq], + exact (rank_quotient_add_rank (linear_map.range (pow_quot_succ_inclusion f p P i))).symm, end -lemma dim_pow_quot [is_domain S] [is_dedekind_domain S] [p.is_maximal] [P.is_prime] +lemma rank_pow_quot [is_domain S] [is_dedekind_domain S] [p.is_maximal] [P.is_prime] (hP0 : P ≠ ⊥) (i : ℕ) (hi : i ≤ e) : module.rank (R ⧸ p) (ideal.map (P^e)^.quotient.mk (P ^ i)) = (e - i) • module.rank (R ⧸ p) (S ⧸ P) := begin refine @nat.decreasing_induction' _ i e (λ j lt_e le_j ih, _) hi _, - { rw [dim_pow_quot_aux f p P _ lt_e, ih, ← succ_nsmul, nat.sub_succ, ← nat.succ_eq_add_one, + { rw [rank_pow_quot_aux f p P _ lt_e, ih, ← succ_nsmul, nat.sub_succ, ← nat.succ_eq_add_one, nat.succ_pred_eq_of_pos (nat.sub_pos_of_lt lt_e)], assumption }, { rw [nat.sub_self, zero_nsmul, map_quotient_self], - exact dim_bot (R ⧸ p) (S ⧸ (P^e)) } + exact rank_bot (R ⧸ p) (S ⧸ (P^e)) } end omit hfp /-- If `p` is a maximal ideal of `R`, `S` extends `R` and `P^e` lies over `p`, then the dimension `[S/(P^e) : R/p]` is equal to `e * [S/P : R/p]`. -/ -lemma dim_prime_pow_ramification_idx [is_domain S] [is_dedekind_domain S] [p.is_maximal] +lemma rank_prime_pow_ramification_idx [is_domain S] [is_dedekind_domain S] [p.is_maximal] [P.is_prime] (hP0 : P ≠ ⊥) (he : e ≠ 0) : module.rank (R ⧸ p) (S ⧸ P^e) = e • @module.rank (R ⧸ p) (S ⧸ P) _ _ (@algebra.to_module _ _ _ _ $ @@quotient.algebra_quotient_of_ramification_idx_ne_zero _ _ _ _ _ ⟨he⟩) := begin letI : ne_zero e := ⟨he⟩, - have := dim_pow_quot f p P hP0 0 (nat.zero_le e), + have := rank_pow_quot f p P hP0 0 (nat.zero_le e), rw [pow_zero, nat.sub_zero, ideal.one_eq_top, ideal.map_top] at this, - exact (dim_top (R ⧸ p) _).symm.trans this + exact (rank_top (R ⧸ p) _).symm.trans this end /-- If `p` is a maximal ideal of `R`, `S` extends `R` and `P^e` lies over `p`, @@ -657,12 +657,12 @@ begin letI : ne_zero e := ⟨he⟩, letI : algebra (R ⧸ p) (S ⧸ P) := quotient.algebra_quotient_of_ramification_idx_ne_zero f p P, letI := ideal.quotient.field p, - have hdim := dim_prime_pow_ramification_idx _ _ _ hP0 he, + have hdim := rank_prime_pow_ramification_idx _ _ _ hP0 he, by_cases hP : finite_dimensional (R ⧸ p) (S ⧸ P), { haveI := hP, haveI := (finite_dimensional_iff_of_rank_eq_nsmul he hdim).mpr hP, refine cardinal.nat_cast_injective _, - rw [finrank_eq_dim, nat.cast_mul, finrank_eq_dim, hdim, nsmul_eq_mul] }, + rw [finrank_eq_rank', nat.cast_mul, finrank_eq_rank', hdim, nsmul_eq_mul] }, have hPe := mt (finite_dimensional_iff_of_rank_eq_nsmul he hdim).mp hP, simp only [finrank_of_infinite_dimensional hP, finrank_of_infinite_dimensional hPe, mul_zero], end From 9c6816cab5872990d450d2c2e7832176167b1c07 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Wed, 5 Apr 2023 13:19:53 +0000 Subject: [PATCH 0780/1291] chore(category_theory/monad/basic): remove some simp lemmas (#18727) This is a backport from mathlib4, where the simpNF (correctly) linter disapproves of these lemmas. Mostly this PR exists to verify that these are not needed in the simp set. https://github.com/leanprover-community/mathlib4/pull/2969 Co-authored-by: Scott Morrison --- src/category_theory/monad/basic.lean | 2 -- 1 file changed, 2 deletions(-) diff --git a/src/category_theory/monad/basic.lean b/src/category_theory/monad/basic.lean index 3ed6b2bcf7ec4..69302b5bd815b 100644 --- a/src/category_theory/monad/basic.lean +++ b/src/category_theory/monad/basic.lean @@ -226,7 +226,6 @@ def monad_to_functor : monad C ⥤ (C ⥤ C) := instance : faithful (monad_to_functor C) := {}. -@[simp] lemma monad_to_functor_map_iso_monad_iso_mk {M N : monad C} (f : (M : C ⥤ C) ≅ N) (f_η f_μ) : (monad_to_functor _).map_iso (monad_iso.mk f f_η f_μ) = f := by { ext, refl } @@ -249,7 +248,6 @@ def comonad_to_functor : comonad C ⥤ (C ⥤ C) := instance : faithful (comonad_to_functor C) := {}. -@[simp] lemma comonad_to_functor_map_iso_comonad_iso_mk {M N : comonad C} (f : (M : C ⥤ C) ≅ N) (f_ε f_δ) : (comonad_to_functor _).map_iso (comonad_iso.mk f f_ε f_δ) = f := by { ext, refl } From 55102fc1d7145d8453f6d35c56d0af6f669f7d12 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 5 Apr 2023 13:19:54 +0000 Subject: [PATCH 0781/1291] chore(linear_algebra/matrix/adjugate): add missing `matrix.of` wrapper (#18736) --- src/linear_algebra/matrix/adjugate.lean | 9 +++++---- 1 file changed, 5 insertions(+), 4 deletions(-) diff --git a/src/linear_algebra/matrix/adjugate.lean b/src/linear_algebra/matrix/adjugate.lean index 9652c44789d95..ba7511931f031 100644 --- a/src/linear_algebra/matrix/adjugate.lean +++ b/src/linear_algebra/matrix/adjugate.lean @@ -169,14 +169,15 @@ These will hold for any matrix over a commutative ring. matrix replacing a column with a basis vector, since it allows us to use facts about the `cramer` map. -/ -def adjugate (A : matrix n n α) : matrix n n α := λ i, cramer Aᵀ (pi.single i 1) +def adjugate (A : matrix n n α) : matrix n n α := +of $ λ i, cramer Aᵀ (pi.single i 1) lemma adjugate_def (A : matrix n n α) : - adjugate A = λ i, cramer Aᵀ (pi.single i 1) := rfl + adjugate A = of (λ i, cramer Aᵀ (pi.single i 1)) := rfl lemma adjugate_apply (A : matrix n n α) (i j : n) : adjugate A i j = (A.update_row j (pi.single i 1)).det := -by { rw adjugate_def, simp only, rw [cramer_apply, update_column_transpose, det_transpose], } +by rw [adjugate_def, of_apply, cramer_apply, update_column_transpose, det_transpose] lemma adjugate_transpose (A : matrix n n α) : (adjugate A)ᵀ = adjugate (Aᵀ) := begin @@ -282,7 +283,7 @@ by { ext, simp [adjugate_def, matrix.one_apply, pi.single_apply, eq_comm] } adjugate (diagonal v) = diagonal (λ i, ∏ j in finset.univ.erase i, v j) := begin ext, - simp only [adjugate_def, cramer_apply, diagonal_transpose], + simp only [adjugate_def, cramer_apply, diagonal_transpose, of_apply], obtain rfl | hij := eq_or_ne i j, { rw [diagonal_apply_eq, diagonal_update_column_single, det_diagonal, prod_update_of_mem (finset.mem_univ _), sdiff_singleton_eq_erase, one_mul] }, From 06a655b5fcfbda03502f9158bbf6c0f1400886f9 Mon Sep 17 00:00:00 2001 From: Yakov Pechersky Date: Wed, 5 Apr 2023 15:13:04 +0000 Subject: [PATCH 0782/1291] feat(data/list/to_finsupp): computable finsupp from list (#15161) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit via `list.nthd` On the way to computable list-based polynomials Co-authored-by: Yakov Pechersky Co-authored-by: Yaël Dillies --- src/data/list/to_finsupp.lean | 145 ++++++++++++++++++++++++++++++++++ src/data/multiset/basic.lean | 2 +- 2 files changed, 146 insertions(+), 1 deletion(-) create mode 100644 src/data/list/to_finsupp.lean diff --git a/src/data/list/to_finsupp.lean b/src/data/list/to_finsupp.lean new file mode 100644 index 0000000000000..3de6de022ed0f --- /dev/null +++ b/src/data/list/to_finsupp.lean @@ -0,0 +1,145 @@ +/- +Copyright (c) 2022 Yakov Pechersky. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yakov Pechersky +-/ + +import data.finsupp.basic + +/-! + +# Lists as finsupp + +# Main definitions + +- `list.to_finsupp`: Interpret a list as a finitely supported function, where the indexing type +is `ℕ`, and the values are either the elements of the list (accessing by indexing) or `0` outside +of the list. + +# Main theorems + +- `list.to_finsupp_eq_sum_map_enum_single`: A `l : list M` over `M` an `add_monoid`, +when interpreted as a finitely supported function, is equal to the sum of `finsupp.single` +produced by mapping over `list.enum l`. + +## Implementation details + +The functions defined here rely on a decidability predicate that each element in the list +can be decidably determined to be not equal to zero or that one can decide one is out of the +bounds of a list. For concretely defined lists that are made up of elements of decidable terms, +this holds. More work will be needed to support lists over non-dec-eq types like `ℝ`, where the +elements are beyond the dec-eq terms of casted values from `ℕ, ℤ, ℚ`. + +-/ + +namespace list + +variables {M : Type*} [has_zero M] (l : list M) + [decidable_pred (λ i, nthd l i 0 ≠ 0)] (n : ℕ) + +/-- Indexing into a `l : list M`, as a finitely-supported function, +where the support are all the indices within the length of the list +that index to a non-zero value. Indices beyond the end of the list are sent to 0. + +This is a computable version of the `finsupp.on_finset` construction. +-/ +def to_finsupp : ℕ →₀ M := +{ to_fun := λ i, nthd l i 0, + support := (finset.range l.length).filter (λ i, nthd l i 0 ≠ 0), + mem_support_to_fun := λ n, begin + simp only [ne.def, finset.mem_filter, finset.mem_range, and_iff_right_iff_imp], + contrapose!, + exact nthd_eq_default _ _ + end } + +@[norm_cast] lemma coe_to_finsupp : (l.to_finsupp : ℕ → M) = λ i, l.nthd i 0 := rfl +@[simp, norm_cast] lemma to_finsupp_apply (i : ℕ) : + (l.to_finsupp : ℕ → M) i = l.nthd i 0 := rfl + +lemma to_finsupp_support : + l.to_finsupp.support = (finset.range l.length).filter (λ i, nthd l i 0 ≠ 0) := +rfl + +lemma to_finsupp_apply_lt (hn : n < l.length) : + l.to_finsupp n = l.nth_le n hn := +nthd_eq_nth_le _ _ _ + +lemma to_finsupp_apply_le (hn : l.length ≤ n) : + l.to_finsupp n = 0 := +nthd_eq_default _ _ hn + +@[simp] lemma to_finsupp_nil [decidable_pred (λ i, nthd ([] : list M) i 0 ≠ 0)] : + to_finsupp ([] : list M) = 0 := +by { ext, simp } + +lemma to_finsupp_singleton (x : M) + [decidable_pred (λ i, nthd [x] i 0 ≠ 0)] : + to_finsupp [x] = finsupp.single 0 x := +begin + ext ⟨_|i⟩; + simp [finsupp.single_apply, (nat.zero_lt_succ _).ne] +end + +@[simp] lemma to_finsupp_cons_apply_zero (x : M) (xs : list M) + [decidable_pred (λ i, nthd (x :: xs) i 0 ≠ 0)] : + (x :: xs).to_finsupp 0 = x := rfl + +@[simp] lemma to_finsupp_cons_apply_succ (x : M) (xs : list M) (n : ℕ) + [decidable_pred (λ i, nthd (x :: xs) i 0 ≠ 0)] + [decidable_pred (λ i, nthd xs i 0 ≠ 0)] : + (x :: xs).to_finsupp n.succ = xs.to_finsupp n := rfl + +lemma to_finsupp_cons_eq_single_add_emb_domain + {R : Type*} [add_zero_class R] (x : R) (xs : list R) + [decidable_pred (λ i, nthd (x :: xs) i 0 ≠ 0)] + [decidable_pred (λ i, nthd xs i 0 ≠ 0)] : + to_finsupp (x :: xs) = finsupp.single 0 x + + (to_finsupp xs).emb_domain ⟨nat.succ, nat.succ_injective⟩ := +begin + ext (_|i), + { simp only [nat.nat_zero_eq_zero, to_finsupp_cons_apply_zero, finsupp.coe_add, + pi.add_apply, finsupp.single_eq_same], + rw finsupp.emb_domain_notin_range, + { exact (add_zero _).symm }, + { simp } }, + { simp only [to_finsupp_cons_apply_succ, finsupp.coe_add, pi.add_apply], + have hi : i.succ = (⟨nat.succ, nat.succ_injective⟩ : ℕ ↪ ℕ) i := rfl, + rw [finsupp.single_apply_eq_zero.mpr, zero_add, hi, finsupp.emb_domain_apply], + simp } +end + +lemma to_finsupp_concat_eq_to_finsupp_add_single + {R : Type*} [add_zero_class R] (x : R) (xs : list R) + [decidable_pred (λ i, nthd (xs ++ [x]) i 0 ≠ 0)] + [decidable_pred (λ i, nthd xs i 0 ≠ 0)] : + to_finsupp (xs ++ [x]) = to_finsupp xs + finsupp.single xs.length x := +begin + ext i, + simp only [finsupp.coe_add, pi.add_apply, finsupp.single_apply], + rcases lt_trichotomy xs.length i with hi|rfl|hi, + { rw [to_finsupp_apply_le _ _ hi.le, to_finsupp_apply_le, + if_neg hi.ne, add_zero], + simpa using nat.succ_le_of_lt hi }, + { rw [to_finsupp_apply_lt, to_finsupp_apply_le _ _ le_rfl, + if_pos rfl, zero_add, nth_le_append_right le_rfl], + { simp }, + { simp } }, + { rw [to_finsupp_apply_lt _ _ hi, to_finsupp_apply_lt, if_neg hi.ne', add_zero, + nth_le_append], + simpa using nat.lt_succ_of_lt hi } +end + +lemma to_finsupp_eq_sum_map_enum_single {R : Type*} [add_monoid R] (l : list R) + [decidable_pred (λ i, nthd l i 0 ≠ 0)] : + to_finsupp l = (l.enum.map (λ (nr : ℕ × R), finsupp.single nr.1 nr.2)).sum := +begin + unfreezingI { induction l using list.reverse_rec_on with xs x IH }, + { convert to_finsupp_nil }, + { simp only [enum_append, map, enum_from_singleton, map_append, sum_append, sum_cons, + sum_nil, add_zero], + classical, + convert to_finsupp_concat_eq_to_finsupp_add_single _ _, + exact IH.symm } +end + +end list diff --git a/src/data/multiset/basic.lean b/src/data/multiset/basic.lean index ae6158c45359a..f58ff719fad96 100644 --- a/src/data/multiset/basic.lean +++ b/src/data/multiset/basic.lean @@ -2279,7 +2279,7 @@ list. -/ def pairwise (r : α → α → Prop) (m : multiset α) : Prop := ∃l:list α, m = l ∧ l.pairwise r -@[simp] lemma pairwise_nil (r : α → α → Prop) : +@[simp] lemma pairwise_zero (r : α → α → Prop) : multiset.pairwise r 0 := ⟨[], rfl, list.pairwise.nil⟩ lemma pairwise_coe_iff {r : α → α → Prop} {l : list α} : From 75be6b616681ab6ca66d798ead117e75cd64f125 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Thu, 6 Apr 2023 01:33:55 +0000 Subject: [PATCH 0783/1291] chore(*): add mathlib4 synchronization comments (#18744) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.char_p.char_and_card` * `category_theory.abelian.basic` * `category_theory.grothendieck` * `category_theory.sites.sheaf` * `combinatorics.simple_graph.adj_matrix` * `combinatorics.simple_graph.inc_matrix` * `computability.ackermann` * `data.string.basic` * `group_theory.perm.cycle.type` * `order.category.Lat` * `order.category.LinOrd` * `order.category.NonemptyFinLinOrd` * `order.category.PartOrd` * `order.category.Preord` * `ring_theory.eisenstein_criterion` * `ring_theory.ideal.quotient_operations` * `topology.algebra.algebra` * `topology.algebra.infinite_sum.module` * `topology.instances.real_vector_space` * `topology.instances.triv_sq_zero_ext` --- src/algebra/char_p/char_and_card.lean | 3 +++ src/category_theory/abelian/basic.lean | 3 +++ src/category_theory/grothendieck.lean | 3 +++ src/category_theory/sites/sheaf.lean | 3 +++ src/combinatorics/simple_graph/adj_matrix.lean | 3 +++ src/combinatorics/simple_graph/inc_matrix.lean | 3 +++ src/computability/ackermann.lean | 3 +++ src/data/string/basic.lean | 3 +++ src/group_theory/perm/cycle/type.lean | 3 +++ src/order/category/Lat.lean | 3 +++ src/order/category/LinOrd.lean | 3 +++ src/order/category/NonemptyFinLinOrd.lean | 3 +++ src/order/category/PartOrd.lean | 3 +++ src/order/category/Preord.lean | 3 +++ src/ring_theory/eisenstein_criterion.lean | 3 +++ src/ring_theory/ideal/quotient_operations.lean | 3 +++ src/topology/algebra/algebra.lean | 3 +++ src/topology/algebra/infinite_sum/module.lean | 5 ++++- src/topology/instances/real_vector_space.lean | 3 +++ src/topology/instances/triv_sq_zero_ext.lean | 3 +++ 20 files changed, 61 insertions(+), 1 deletion(-) diff --git a/src/algebra/char_p/char_and_card.lean b/src/algebra/char_p/char_and_card.lean index 10d61ec6c3708..52b303e3667de 100644 --- a/src/algebra/char_p/char_and_card.lean +++ b/src/algebra/char_p/char_and_card.lean @@ -9,6 +9,9 @@ import group_theory.perm.cycle.type /-! # Characteristic and cardinality +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We prove some results relating characteristic and cardinality of finite rings ## Tags diff --git a/src/category_theory/abelian/basic.lean b/src/category_theory/abelian/basic.lean index 160b7c3b88b08..bfbd35d438ca0 100644 --- a/src/category_theory/abelian/basic.lean +++ b/src/category_theory/abelian/basic.lean @@ -13,6 +13,9 @@ import category_theory.abelian.non_preadditive /-! # Abelian categories +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains the definition and basic properties of abelian categories. There are many definitions of abelian category. Our definition is as follows: diff --git a/src/category_theory/grothendieck.lean b/src/category_theory/grothendieck.lean index c304cd2f89a7f..35512dd11a57a 100644 --- a/src/category_theory/grothendieck.lean +++ b/src/category_theory/grothendieck.lean @@ -9,6 +9,9 @@ import category_theory.elements /-! # The Grothendieck construction +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given a functor `F : C ⥤ Cat`, the objects of `grothendieck F` consist of dependent pairs `(b, f)`, where `b : C` and `f : F.obj c`, and a morphism `(b, f) ⟶ (b', f')` is a pair `β : b ⟶ b'` in `C`, and diff --git a/src/category_theory/sites/sheaf.lean b/src/category_theory/sites/sheaf.lean index e9ad13949f38a..149a0b9049837 100644 --- a/src/category_theory/sites/sheaf.lean +++ b/src/category_theory/sites/sheaf.lean @@ -13,6 +13,9 @@ import category_theory.sites.sheaf_of_types /-! # Sheaves taking values in a category +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + If C is a category with a Grothendieck topology, we define the notion of a sheaf taking values in an arbitrary category `A`. We follow the definition in https://stacks.math.columbia.edu/tag/00VR, noting that the presheaf of sets "defined above" can be seen in the comments between tags 00VQ and diff --git a/src/combinatorics/simple_graph/adj_matrix.lean b/src/combinatorics/simple_graph/adj_matrix.lean index fbb570bf19eb3..fa1cde99ecd1d 100644 --- a/src/combinatorics/simple_graph/adj_matrix.lean +++ b/src/combinatorics/simple_graph/adj_matrix.lean @@ -11,6 +11,9 @@ import linear_algebra.matrix.symmetric /-! # Adjacency Matrices +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This module defines the adjacency matrix of a graph, and provides theorems connecting graph properties to computational properties of the matrix. diff --git a/src/combinatorics/simple_graph/inc_matrix.lean b/src/combinatorics/simple_graph/inc_matrix.lean index a7ef63ba4b7f7..46829ff8eceb3 100644 --- a/src/combinatorics/simple_graph/inc_matrix.lean +++ b/src/combinatorics/simple_graph/inc_matrix.lean @@ -9,6 +9,9 @@ import data.matrix.basic /-! # Incidence matrix of a simple graph +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the unoriented incidence matrix of a simple graph. ## Main definitions diff --git a/src/computability/ackermann.lean b/src/computability/ackermann.lean index 8556260af6365..37a83d6c913b0 100644 --- a/src/computability/ackermann.lean +++ b/src/computability/ackermann.lean @@ -9,6 +9,9 @@ import tactic.linarith /-! # Ackermann function +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we define the two-argument Ackermann function `ack`. Despite having a recursive definition, we show that this isn't a primitive recursive function. diff --git a/src/data/string/basic.lean b/src/data/string/basic.lean index 594a47f40015b..11b9289d66277 100644 --- a/src/data/string/basic.lean +++ b/src/data/string/basic.lean @@ -9,6 +9,9 @@ import data.char /-! # Strings +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Supplementary theorems about the `string` type. -/ diff --git a/src/group_theory/perm/cycle/type.lean b/src/group_theory/perm/cycle/type.lean index 0a8dcbff71766..5dfacba2f3c54 100644 --- a/src/group_theory/perm/cycle/type.lean +++ b/src/group_theory/perm/cycle/type.lean @@ -14,6 +14,9 @@ import tactic.linarith /-! # Cycle Types +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define the cycle type of a permutation. ## Main definitions diff --git a/src/order/category/Lat.lean b/src/order/category/Lat.lean index 85a51ca335bc8..5fe82c5f481fe 100644 --- a/src/order/category/Lat.lean +++ b/src/order/category/Lat.lean @@ -9,6 +9,9 @@ import order.hom.lattice /-! # The category of lattices +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This defines `Lat`, the category of lattices. Note that `Lat` doesn't correspond to the literature definition of [`Lat`] diff --git a/src/order/category/LinOrd.lean b/src/order/category/LinOrd.lean index f3eefd21eeba9..1c9d29bf5d82a 100644 --- a/src/order/category/LinOrd.lean +++ b/src/order/category/LinOrd.lean @@ -9,6 +9,9 @@ import order.category.Lat /-! # Category of linear orders +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This defines `LinOrd`, the category of linear orders with monotone maps. -/ diff --git a/src/order/category/NonemptyFinLinOrd.lean b/src/order/category/NonemptyFinLinOrd.lean index af7ab122b152c..db8e1d206bc0d 100644 --- a/src/order/category/NonemptyFinLinOrd.lean +++ b/src/order/category/NonemptyFinLinOrd.lean @@ -12,6 +12,9 @@ import category_theory.limits.shapes.regular_mono /-! # Nonempty finite linear orders +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This defines `NonemptyFinLinOrd`, the category of nonempty finite linear orders with monotone maps. This is the index category for simplicial objects. -/ diff --git a/src/order/category/PartOrd.lean b/src/order/category/PartOrd.lean index b4baeee3989f7..f4fdaa6bedf90 100644 --- a/src/order/category/PartOrd.lean +++ b/src/order/category/PartOrd.lean @@ -9,6 +9,9 @@ import order.category.Preord /-! # Category of partial orders +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This defines `PartOrd`, the category of partial orders with monotone maps. -/ diff --git a/src/order/category/Preord.lean b/src/order/category/Preord.lean index ee40020817599..1ca4aea91d877 100644 --- a/src/order/category/Preord.lean +++ b/src/order/category/Preord.lean @@ -11,6 +11,9 @@ import order.hom.basic /-! # Category of preorders +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This defines `Preord`, the category of preorders with monotone maps. -/ diff --git a/src/ring_theory/eisenstein_criterion.lean b/src/ring_theory/eisenstein_criterion.lean index fe9dc886ef0b0..78fc40f281f19 100644 --- a/src/ring_theory/eisenstein_criterion.lean +++ b/src/ring_theory/eisenstein_criterion.lean @@ -10,6 +10,9 @@ import ring_theory.ideal.quotient_operations /-! # Eisenstein's criterion +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A proof of a slight generalisation of Eisenstein's criterion for the irreducibility of a polynomial over an integral domain. -/ diff --git a/src/ring_theory/ideal/quotient_operations.lean b/src/ring_theory/ideal/quotient_operations.lean index 4bbea539b3a0e..8eee8b5ab2641 100644 --- a/src/ring_theory/ideal/quotient_operations.lean +++ b/src/ring_theory/ideal/quotient_operations.lean @@ -7,6 +7,9 @@ import ring_theory.ideal.operations import ring_theory.ideal.quotient /-! # More operations on modules and ideals related to quotients + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ universes u v w diff --git a/src/topology/algebra/algebra.lean b/src/topology/algebra/algebra.lean index 7057336c41ad8..539fcb7629988 100644 --- a/src/topology/algebra/algebra.lean +++ b/src/topology/algebra/algebra.lean @@ -10,6 +10,9 @@ import ring_theory.adjoin.basic /-! # Topological (sub)algebras +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A topological algebra over a topological semiring `R` is a topological semiring with a compatible continuous scalar multiplication by elements of `R`. We reuse typeclass `has_continuous_smul` for topological algebras. diff --git a/src/topology/algebra/infinite_sum/module.lean b/src/topology/algebra/infinite_sum/module.lean index 5ec77641df867..ad9f0b35f7f74 100644 --- a/src/topology/algebra/infinite_sum/module.lean +++ b/src/topology/algebra/infinite_sum/module.lean @@ -6,7 +6,10 @@ Authors: Heather Macbeth, Yury Kudryashov, Frédéric Dupuis import topology.algebra.infinite_sum.basic import topology.algebra.module.basic -/-! # Infinite sums in topological vector spaces -/ +/-! # Infinite sums in topological vector spaces + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4.-/ variables {ι R R₂ M M₂ : Type*} diff --git a/src/topology/instances/real_vector_space.lean b/src/topology/instances/real_vector_space.lean index 05e34574b12b7..6f7575e36dc96 100644 --- a/src/topology/instances/real_vector_space.lean +++ b/src/topology/instances/real_vector_space.lean @@ -9,6 +9,9 @@ import topology.instances.rat /-! # Continuous additive maps are `ℝ`-linear +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove that a continuous map `f : E →+ F` between two topological vector spaces over `ℝ` is `ℝ`-linear -/ diff --git a/src/topology/instances/triv_sq_zero_ext.lean b/src/topology/instances/triv_sq_zero_ext.lean index c791c70b0f1dd..1903be2866457 100644 --- a/src/topology/instances/triv_sq_zero_ext.lean +++ b/src/topology/instances/triv_sq_zero_ext.lean @@ -10,6 +10,9 @@ import topology.algebra.module.basic /-! # Topology on `triv_sq_zero_ext R M` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The type `triv_sq_zero_ext R M` inherits the topology from `R × M`. Note that this is not the topology induced by the seminorm on the dual numbers suggested by From 5aa3c1de9f3c642eac76e11071c852766f220fd0 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 6 Apr 2023 08:23:25 +0000 Subject: [PATCH 0784/1291] chore(linear_algebra/dimension): deduplicate lemmas (#18743) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We have some lemmas in the `module.free` namespace which duplicate lemmas in the root namespace. This moves all the remaining `rank` lemmas in this namespace into the root namespace, and cleans up the overlapping lemmas this creates. The changes are: * `module.free.rank_eq_card_choose_basis_index`: unchanged but moved to an earlier file * `rank_prod` → `rank_prod'` (the non-universe polymorphic version) * `module.free.rank_prod` → `rank_prod` * none → `rank_finsupp` (new lemma) * `finsupp.rank_eq` → `rank_finsupp'` (the non-universe polymorphic version) * `module.free.rank_finsupp` → `rank_finsupp_self` * `module.free.rank_finsupp'` → `rank_finsupp_self'` (the non-universe polymorphic version) * `module.free.rank_pi_finite` → `rank_pi` (these were duplicates) * For everything else, `module.free.rank_*` → `rank_*` --- src/field_theory/finite/polynomial.lean | 3 +- src/linear_algebra/dimension.lean | 60 ++++++++++--------- .../free_module/finite/rank.lean | 4 +- src/linear_algebra/free_module/rank.lean | 52 ++++++---------- 4 files changed, 53 insertions(+), 66 deletions(-) diff --git a/src/field_theory/finite/polynomial.lean b/src/field_theory/finite/polynomial.lean index b3b4000b72398..97d574232fdc3 100644 --- a/src/field_theory/finite/polynomial.lean +++ b/src/field_theory/finite/polynomial.lean @@ -182,8 +182,7 @@ calc module.rank K (R σ K) = module.rank K (↥{s : σ →₀ ℕ | ∀ (n : σ), s n ≤ fintype.card K - 1} →₀ K) : linear_equiv.rank_eq (finsupp.supported_equiv_finsupp {s : σ →₀ ℕ | ∀n:σ, s n ≤ fintype.card K - 1 }) - ... = #{s : σ →₀ ℕ | ∀ (n : σ), s n ≤ fintype.card K - 1} : - by rw [finsupp.rank_eq, rank_self, mul_one] + ... = #{s : σ →₀ ℕ | ∀ (n : σ), s n ≤ fintype.card K - 1} : by rw [rank_finsupp_self'] ... = #{s : σ → ℕ | ∀ (n : σ), s n < fintype.card K } : begin refine quotient.sound ⟨equiv.subtype_equiv finsupp.equiv_fun_on_finite $ assume f, _⟩, diff --git a/src/linear_algebra/dimension.lean b/src/linear_algebra/dimension.lean index 61da11984d18c..4461772fbe840 100644 --- a/src/linear_algebra/dimension.lean +++ b/src/linear_algebra/dimension.lean @@ -942,6 +942,19 @@ variables [add_comm_group V'] [module K V'] [module.free K V'] variables [add_comm_group V₁] [module K V₁] [module.free K V₁] variables {K V} + +namespace module.free +variables (K V) + +/-- The rank of a free module `M` over `R` is the cardinality of `choose_basis_index R M`. -/ +lemma rank_eq_card_choose_basis_index : module.rank K V = #(choose_basis_index K V) := +(choose_basis K V).mk_eq_rank''.symm + +end module.free + +open module.free +open cardinal + /-- Two vector spaces are isomorphic if they have the same dimension. -/ theorem nonempty_linear_equiv_of_lift_rank_eq (cond : cardinal.lift.{v'} (module.rank K V) = cardinal.lift.{v} (module.rank K V')) : @@ -987,35 +1000,31 @@ theorem linear_equiv.nonempty_equiv_iff_rank_eq : nonempty (V ≃ₗ[K] V₁) ↔ module.rank K V = module.rank K V₁ := ⟨λ ⟨h⟩, linear_equiv.rank_eq h, λ h, nonempty_linear_equiv_of_rank_eq h⟩ -theorem rank_prod : module.rank K (V × V₁) = module.rank K V + module.rank K V₁ := -begin - obtain ⟨⟨_, b⟩⟩ := module.free.exists_basis K V, - obtain ⟨⟨_, c⟩⟩ := module.free.exists_basis K V₁, - rw [← cardinal.lift_inj, - ← (basis.prod b c).mk_eq_rank, - cardinal.lift_add, ← cardinal.mk_ulift, - ← b.mk_eq_rank, ← c.mk_eq_rank, - ← cardinal.mk_ulift, ← cardinal.mk_ulift, - cardinal.add_def (ulift _)], - exact cardinal.lift_inj.1 (cardinal.lift_mk_eq.2 - ⟨equiv.ulift.trans (equiv.sum_congr equiv.ulift equiv.ulift).symm ⟩), -end +/-- The rank of `M × N` is `(module.rank R M).lift + (module.rank R N).lift`. -/ +@[simp] lemma rank_prod : + module.rank K (V × V') = + cardinal.lift.{v'} (module.rank K V) + cardinal.lift.{v v'} (module.rank K V') := +by simpa [rank_eq_card_choose_basis_index K V, rank_eq_card_choose_basis_index K V', + lift_umax, lift_umax'] using ((choose_basis K V).prod (choose_basis K V')).mk_eq_rank.symm + +/-- If `M` and `N` lie in the same universe, the rank of `M × N` is + `(module.rank R M) + (module.rank R N)`. -/ +theorem rank_prod' : module.rank K (V × V₁) = module.rank K V + module.rank K V₁ := +by simp section fintype variables [∀i, add_comm_group (φ i)] [∀i, module K (φ i)] [∀i, module.free K (φ i)] open linear_map -lemma rank_pi [finite η] : +/-- The rank of a finite product is the sum of the ranks. -/ +@[simp] lemma rank_pi [finite η] : module.rank K (Πi, φ i) = cardinal.sum (λi, module.rank K (φ i)) := begin - haveI := nontrivial_of_invariant_basis_number K, casesI nonempty_fintype η, - let b := λ i, (module.free.exists_basis K (φ i)).some.2, - let this : basis (Σ j, _) K (Π j, φ j) := pi.basis b, - rw [← cardinal.lift_inj, ← this.mk_eq_rank], - simp_rw [cardinal.mk_sigma, cardinal.lift_sum, ←(b _).mk_range_eq_rank, - cardinal.mk_range_eq _ (b _).injective], + let B := λ i, choose_basis K (φ i), + let b : basis _ K (Π i, φ i) := pi.basis (λ i, B i), + simp [← b.mk_eq_rank'', λ i, (B i).mk_eq_rank''], end variable [fintype η] @@ -1038,13 +1047,6 @@ by simp [rank_fun'] end fintype -lemma finsupp.rank_eq {ι : Type v} : module.rank K (ι →₀ V) = #ι * module.rank K V := -begin - obtain ⟨⟨_, bs⟩⟩ := module.free.exists_basis K V, - rw [← bs.mk_eq_rank'', ← (finsupp.basis (λa:ι, bs)).mk_eq_rank'', - cardinal.mk_sigma, cardinal.sum_const'] -end - -- TODO: merge with the `finrank` content /-- An `n`-dimensional `K`-vector space is equivalent to `fin n → K`. -/ def fin_dim_vectorspace_equiv (n : ℕ) @@ -1092,7 +1094,7 @@ calc module.rank K (span K (↑s : set V)) ≤ #(↑s : set V) : rank_span_le theorem rank_quotient_add_rank (p : submodule K V) : module.rank K (V ⧸ p) + module.rank K p = module.rank K V := -by classical; exact let ⟨f⟩ := quotient_prod_linear_equiv p in rank_prod.symm.trans f.rank_eq +by classical; exact let ⟨f⟩ := quotient_prod_linear_equiv p in rank_prod'.symm.trans f.rank_eq /-- rank-nullity theorem -/ theorem rank_range_add_rank_ker (f : V →ₗ[K] V₁) : @@ -1122,7 +1124,7 @@ lemma rank_add_rank_split have hf : surjective (coprod db eb), by rwa [←range_eq_top, range_coprod, eq_top_iff], begin - conv {to_rhs, rw [← rank_prod, rank_eq_of_surjective _ hf] }, + conv {to_rhs, rw [← rank_prod', rank_eq_of_surjective _ hf] }, congr' 1, apply linear_equiv.rank_eq, refine linear_equiv.of_bijective _ ⟨_, _⟩, diff --git a/src/linear_algebra/free_module/finite/rank.lean b/src/linear_algebra/free_module/finite/rank.lean index 9e801be89d6fd..dfcfef1346576 100644 --- a/src/linear_algebra/free_module/finite/rank.lean +++ b/src/linear_algebra/free_module/finite/rank.lean @@ -57,7 +57,7 @@ end /-- The finrank of `(ι →₀ R)` is `fintype.card ι`. -/ @[simp] lemma finrank_finsupp {ι : Type v} [fintype ι] : finrank R (ι →₀ R) = card ι := -by { rw [finrank, rank_finsupp, ← mk_to_nat_eq_card, to_nat_lift] } +by { rw [finrank, rank_finsupp_self, ← mk_to_nat_eq_card, to_nat_lift] } /-- The finrank of `(ι → R)` is `fintype.card ι`. -/ lemma finrank_pi {ι : Type v} [fintype ι] : finrank R (ι → R) = card ι := @@ -84,7 +84,7 @@ lemma finrank_pi_fintype {ι : Type v} [fintype ι] {M : ι → Type w} [Π (i : ι), module.finite R (M i)] : finrank R (Π i, M i) = ∑ i, finrank R (M i) := begin letI := nontrivial_of_invariant_basis_number R, - simp only [finrank, λ i, rank_eq_card_choose_basis_index R (M i), rank_pi_finite, + simp only [finrank, λ i, rank_eq_card_choose_basis_index R (M i), rank_pi, ← mk_sigma, mk_to_nat_eq_card, card_sigma], end diff --git a/src/linear_algebra/free_module/rank.lean b/src/linear_algebra/free_module/rank.lean index 5eb9e26f3617c..5fed1512b65cf 100644 --- a/src/linear_algebra/free_module/rank.lean +++ b/src/linear_algebra/free_module/rank.lean @@ -5,14 +5,12 @@ Authors: Riccardo Brasca -/ import linear_algebra.dimension -import linear_algebra.free_module.basic -import linear_algebra.invariant_basis_number /-! -# Rank of free modules +# Extra results about `module.rank` -This is a basic API for the rank of free modules. +This file contains some extra results not in `linear_algebra.dimension`. -/ @@ -24,36 +22,31 @@ open_locale tensor_product direct_sum big_operators cardinal open cardinal -namespace module.free - section ring variables [ring R] [strong_rank_condition R] variables [add_comm_group M] [module R M] [module.free R M] variables [add_comm_group N] [module R N] [module.free R N] -/-- The rank of a free module `M` over `R` is the cardinality of `choose_basis_index R M`. -/ -lemma rank_eq_card_choose_basis_index : module.rank R M = #(choose_basis_index R M) := -(choose_basis R M).mk_eq_rank''.symm +open module.free -/-- The rank of `(ι →₀ R)` is `(# ι).lift`. -/ -@[simp] lemma rank_finsupp {ι : Type v} : module.rank R (ι →₀ R) = (# ι).lift := -by simpa [lift_id', lift_umax] using - (basis.of_repr (linear_equiv.refl _ (ι →₀ R))).mk_eq_rank.symm +@[simp] lemma rank_finsupp (ι : Type w) : + module.rank R (ι →₀ M) = cardinal.lift.{v} #ι * cardinal.lift.{w} (module.rank R M) := +begin + obtain ⟨⟨_, bs⟩⟩ := module.free.exists_basis R M, + rw [← bs.mk_eq_rank'', ← (finsupp.basis (λa:ι, bs)).mk_eq_rank'', + cardinal.mk_sigma, cardinal.sum_const] +end -/-- If `R` and `ι` lie in the same universe, the rank of `(ι →₀ R)` is `# ι`. -/ -lemma rank_finsupp' {ι : Type u} : module.rank R (ι →₀ R) = # ι := by simp +lemma rank_finsupp' (ι : Type v) : module.rank R (ι →₀ M) = #ι * module.rank R M := +by simp [rank_finsupp] -/-- The rank of `M × N` is `(module.rank R M).lift + (module.rank R N).lift`. -/ -@[simp] lemma rank_prod : - module.rank R (M × N) = lift.{w v} (module.rank R M) + lift.{v w} (module.rank R N) := -by simpa [rank_eq_card_choose_basis_index R M, rank_eq_card_choose_basis_index R N, - lift_umax, lift_umax'] using ((choose_basis R M).prod (choose_basis R N)).mk_eq_rank.symm +/-- The rank of `(ι →₀ R)` is `(# ι).lift`. -/ +@[simp] lemma rank_finsupp_self (ι : Type w) : module.rank R (ι →₀ R) = (# ι).lift := +by simp [rank_finsupp] -/-- If `M` and `N` lie in the same universe, the rank of `M × N` is - `(module.rank R M) + (module.rank R N)`. -/ -lemma rank_prod' (N : Type v) [add_comm_group N] [module R N] [module.free R N] : - module.rank R (M × N) = (module.rank R M) + (module.rank R N) := by simp +/-- If `R` and `ι` lie in the same universe, the rank of `(ι →₀ R)` is `# ι`. -/ +lemma rank_finsupp_self' {ι : Type u} : module.rank R (ι →₀ R) = # ι := by simp /-- The rank of the direct sum is the sum of the ranks. -/ @[simp] lemma rank_direct_sum {ι : Type v} (M : ι → Type w) [Π (i : ι), add_comm_group (M i)] @@ -65,13 +58,6 @@ begin simp [← b.mk_eq_rank'', λ i, (B i).mk_eq_rank''], end -/-- The rank of a finite product is the sum of the ranks. -/ -@[simp] lemma rank_pi_finite {ι : Type v} [finite ι] {M : ι → Type w} - [Π (i : ι), add_comm_group (M i)] [Π (i : ι), module R (M i)] [Π (i : ι), module.free R (M i)] : - module.rank R (Π i, M i) = cardinal.sum (λ i, module.rank R (M i)) := -by { casesI nonempty_fintype ι, - rw [←(direct_sum.linear_equiv_fun_on_fintype _ _ M).rank_eq, rank_direct_sum] } - /-- If `m` and `n` are `fintype`, the rank of `m × n` matrices is `(# m).lift * (# n).lift`. -/ @[simp] lemma rank_matrix (m : Type v) (n : Type w) [finite m] [finite n] : module.rank R (matrix m n R) = (lift.{(max v w u) v} (# m)) * (lift.{(max v w u) w} (# n)) := @@ -102,6 +88,8 @@ variables [comm_ring R] [strong_rank_condition R] variables [add_comm_group M] [module R M] [module.free R M] variables [add_comm_group N] [module R N] [module.free R N] +open module.free + /-- The rank of `M ⊗[R] N` is `(module.rank R M).lift * (module.rank R N).lift`. -/ @[simp] lemma rank_tensor_product : module.rank R (M ⊗[R] N) = lift.{w v} (module.rank R M) * lift.{v w} (module.rank R N) := @@ -121,5 +109,3 @@ lemma rank_tensor_product' (N : Type v) [add_comm_group N] [module R N] [module. module.rank R (M ⊗[R] N) = (module.rank R M) * (module.rank R N) := by simp end comm_ring - -end module.free From 60da01b41bbe4206f05d34fd70c8dd7498717a30 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= Date: Thu, 6 Apr 2023 10:22:10 +0000 Subject: [PATCH 0785/1291] feat(number_theory/number_field/canonical_embedding): add canonical embedding (#17783) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR defines the canonical embedding of a number field $K$ of signature $(r_1, r_2)$ into $\mathbb{R}^{r_1} × \mathbb{C}^{r_2}$ and prove various results about this embedding, including the fact that the image of the ring of integers of $K$ is discrete. Co-authored-by: Alex J Best Co-authored-by: Xavier-François Roblot <46200072+xroblot@users.noreply.github.com> --- .../number_field/canonical_embedding.lean | 166 ++++++++++++++++++ .../number_field/embeddings.lean | 14 +- 2 files changed, 178 insertions(+), 2 deletions(-) create mode 100644 src/number_theory/number_field/canonical_embedding.lean diff --git a/src/number_theory/number_field/canonical_embedding.lean b/src/number_theory/number_field/canonical_embedding.lean new file mode 100644 index 0000000000000..bdcc13bebd7ef --- /dev/null +++ b/src/number_theory/number_field/canonical_embedding.lean @@ -0,0 +1,166 @@ +/- +Copyright (c) 2022 Xavier Roblot. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Xavier Roblot +-/ +import number_theory.number_field.embeddings + +/-! +# Canonical embedding of a number field + +The canonical embedding of a number field `K` of signature `(r₁, r₂)` is the ring homomorphism +`K →+* ℝ^r₁ × ℂ^r₂` that sends `x ∈ K` to `(φ_₁(x),...,φ_r₁(x)) × (ψ_₁(x),..., ψ_r₂(x))` where +`φ_₁,...,φ_r₁` are its real embeddings and `ψ_₁,..., ψ_r₂` are its complex embeddings (up to +complex conjugation). + +## Main definitions and results + +* `number_field.canonical_embedding.ring_of_integers.inter_ball_finite`: the intersection of the +image of the ring of integers by the canonical embedding and any ball centered at `0` of finite +radius is finite. + +## Tags + +number field, infinite places +-/ + +noncomputable theory + +open function finite_dimensional finset fintype number_field number_field.infinite_place metric +module +open_locale classical number_field + +variables (K : Type*) [field K] + +namespace number_field.canonical_embedding + +-- The ambient space `ℝ^r₁ × ℂ^r₂` with `(r₁, r₂)` the signature of `K`. +localized "notation `E` := + ({w : infinite_place K // is_real w} → ℝ) × ({w : infinite_place K // is_complex w} → ℂ)" + in canonical_embedding + +lemma space_rank [number_field K] : + finrank ℝ E = finrank ℚ K := +begin + haveI : module.free ℝ ℂ := infer_instance, + rw [finrank_prod, finrank_pi, finrank_pi_fintype, complex.finrank_real_complex, + finset.sum_const, finset.card_univ, ← card_real_embeddings, algebra.id.smul_eq_mul, mul_comm, + ← card_complex_embeddings, ← number_field.embeddings.card K ℂ, fintype.card_subtype_compl, + nat.add_sub_of_le (fintype.card_subtype_le _)], +end + +lemma non_trivial_space [number_field K] : nontrivial E := +begin + obtain ⟨w⟩ := infinite_place.nonempty K, + obtain hw | hw := w.is_real_or_is_complex, + { haveI : nonempty {w : infinite_place K // is_real w} := ⟨⟨w, hw⟩⟩, + exact nontrivial_prod_left }, + { haveI : nonempty {w : infinite_place K // is_complex w} := ⟨⟨w, hw⟩⟩, + exact nontrivial_prod_right } +end + +/-- The canonical embedding of a number field `K` of signature `(r₁, r₂)` into `ℝ^r₁ × ℂ^r₂`. -/ +def _root_.number_field.canonical_embedding : K →+* E := +ring_hom.prod (pi.ring_hom (λ w, w.prop.embedding)) (pi.ring_hom (λ w, w.val.embedding)) + +lemma _root_.number_field.canonical_embedding_injective [number_field K] : + injective (number_field.canonical_embedding K) := + @ring_hom.injective _ _ _ _ (non_trivial_space K) _ + +open number_field + +variable {K} + +@[simp] +lemma apply_at_real_infinite_place (w : {w : infinite_place K // is_real w}) (x : K) : + (number_field.canonical_embedding K x).1 w = w.prop.embedding x := +by simp only [canonical_embedding, ring_hom.prod_apply, pi.ring_hom_apply] + +@[simp] +lemma apply_at_complex_infinite_place (w : { w : infinite_place K // is_complex w}) (x : K) : + (number_field.canonical_embedding K x).2 w = embedding w.val x := +by simp only [canonical_embedding, ring_hom.prod_apply, pi.ring_hom_apply] + +lemma nnnorm_eq [number_field K] (x : K) : + ‖canonical_embedding K x‖₊ = finset.univ.sup (λ w : infinite_place K, ⟨w x, map_nonneg w x⟩) := +begin + rw [prod.nnnorm_def', pi.nnnorm_def, pi.nnnorm_def], + rw ( _ : finset.univ = {w : infinite_place K | is_real w}.to_finset + ∪ {w : infinite_place K | is_complex w}.to_finset), + { rw [finset.sup_union, sup_eq_max], + refine congr_arg2 _ _ _, + { convert (finset.univ.sup_map (function.embedding.subtype (λ w : infinite_place K, is_real w)) + (λ w, (⟨w x, map_nonneg w x⟩ : nnreal))).symm using 2, + ext w, + simp only [apply_at_real_infinite_place, coe_nnnorm, real.norm_eq_abs, + function.embedding.coe_subtype, subtype.coe_mk, is_real.abs_embedding_apply], }, + { convert (finset.univ.sup_map (function.embedding.subtype (λ w : infinite_place K, + is_complex w)) (λ w, (⟨w x, map_nonneg w x⟩ : nnreal))).symm using 2, + ext w, + simp only [apply_at_complex_infinite_place, subtype.val_eq_coe, coe_nnnorm, + complex.norm_eq_abs, function.embedding.coe_subtype, subtype.coe_mk, abs_embedding], }}, + { ext w, + simp only [w.is_real_or_is_complex, set.mem_set_of_eq, finset.mem_union, set.mem_to_finset, + finset.mem_univ], }, +end + +lemma norm_le_iff [number_field K] (x : K) (r : ℝ) : + ‖canonical_embedding K x‖ ≤ r ↔ ∀ w : infinite_place K, w x ≤ r := +begin + obtain hr | hr := lt_or_le r 0, + { obtain ⟨w⟩ := infinite_place.nonempty K, + exact iff_of_false (hr.trans_le $ norm_nonneg _).not_le + (λ h, hr.not_le $ (map_nonneg w _).trans $ h _) }, + { lift r to nnreal using hr, + simp_rw [← coe_nnnorm, nnnorm_eq, nnreal.coe_le_coe, finset.sup_le_iff, finset.mem_univ, + forall_true_left, ←nnreal.coe_le_coe, subtype.coe_mk] } +end + +variables (K) + +/-- The image of `𝓞 K` as a subring of `ℝ^r₁ × ℂ^r₂`. -/ +def integer_lattice : subring E := +(ring_hom.range (algebra_map (𝓞 K) K)).map (canonical_embedding K) + +/-- The linear equiv between `𝓞 K` and the integer lattice. -/ +def equiv_integer_lattice [number_field K] : + 𝓞 K ≃ₗ[ℤ] integer_lattice K := +linear_equiv.of_bijective + { to_fun := λ x, ⟨canonical_embedding K (algebra_map (𝓞 K) K x), algebra_map (𝓞 K) K x, + by simp only [subring.mem_carrier, ring_hom.mem_range, exists_apply_eq_apply], rfl⟩, + map_add' := λ x y, by simpa only [map_add], + map_smul' := λ c x, by simpa only [zsmul_eq_mul, map_mul, map_int_cast] } + begin + refine ⟨λ _ _ h, _, λ ⟨_, _, ⟨a, rfl⟩, rfl⟩, ⟨a, rfl⟩⟩, + rw [linear_map.coe_mk, subtype.mk_eq_mk] at h, + exact is_fraction_ring.injective (𝓞 K) K (canonical_embedding_injective K h), + end + +lemma integer_lattice.inter_ball_finite [number_field K] (r : ℝ) : + ((integer_lattice K : set E) ∩ (closed_ball 0 r)).finite := +begin + obtain hr | hr := lt_or_le r 0, + { simp [closed_ball_eq_empty.2 hr] }, + have heq : + ∀ x, canonical_embedding K x ∈ closed_ball (0 : E) r ↔ ∀ φ : K →+* ℂ, ‖φ x‖ ≤ r, + { simp only [← place_apply, ← infinite_place.coe_mk, mem_closed_ball_zero_iff, norm_le_iff], + exact λ x, le_iff_le x r, }, + convert (embeddings.finite_of_norm_le K ℂ r).image (canonical_embedding K), + ext, split, + { rintro ⟨⟨_, ⟨x, rfl⟩, rfl⟩, hx2⟩, + exact ⟨x, ⟨set_like.coe_mem x, (heq x).mp hx2⟩, rfl⟩, }, + { rintro ⟨x, ⟨hx1, hx2⟩, rfl⟩, + exact ⟨⟨x, ⟨⟨x, hx1⟩, rfl⟩, rfl⟩, (heq x).mpr hx2⟩, } +end + +instance [number_field K] : countable (integer_lattice K) := +begin + have : (⋃ n : ℕ, ((integer_lattice K : set E) ∩ (closed_ball 0 n))).countable, + { exact set.countable_Union (λ n, (integer_lattice.inter_ball_finite K n).countable) }, + refine (this.mono _).to_subtype, + rintro _ ⟨x, hx, rfl⟩, + rw set.mem_Union, + exact ⟨⌈‖canonical_embedding K x‖⌉₊, ⟨x, hx, rfl⟩, mem_closed_ball_zero_iff.2 (nat.le_ceil _)⟩, +end + +end number_field.canonical_embedding diff --git a/src/number_theory/number_field/embeddings.lean b/src/number_theory/number_field/embeddings.lean index cf7b9af532659..ed90ed59b4282 100644 --- a/src/number_theory/number_field/embeddings.lean +++ b/src/number_theory/number_field/embeddings.lean @@ -338,10 +338,15 @@ begin { exact λ h, ⟨embedding w, h, mk_embedding w⟩, }, end -lemma not_is_real_iff_is_complex {w : infinite_place K} : - ¬ is_real w ↔ is_complex w := +@[simp] lemma not_is_real_iff_is_complex {w : infinite_place K} : ¬ is_real w ↔ is_complex w := by rw [is_complex_iff, is_real_iff] +@[simp] lemma not_is_complex_iff_is_real {w : infinite_place K} : ¬ is_complex w ↔ is_real w := +by rw [←not_is_real_iff_is_complex, not_not] + +lemma is_real_or_is_complex (w : infinite_place K) : is_real w ∨ is_complex w := +by { rw ←not_is_real_iff_is_complex, exact em _ } + /-- For `w` a real infinite place, return the corresponding embedding as a morphism `K →+* ℝ`. -/ noncomputable def is_real.embedding {w : infinite_place K} (hw : is_real w) : K →+* ℝ := (is_real_iff.mp hw).embedding @@ -354,6 +359,11 @@ begin exact congr_fun (congr_arg coe_fn (mk_embedding w)) x, end +@[simp] +lemma is_real.abs_embedding_apply {w : infinite_place K} (hw : is_real w) (x : K) : + |is_real.embedding hw x| = w x := +by { rw ← is_real.place_embedding_apply hw x, congr, } + variable (K) /-- The map from real embeddings to real infinite places as an equiv -/ From 67e606eaea14c7854bdc556bd53d98aefdf76ec0 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 6 Apr 2023 12:45:35 +0000 Subject: [PATCH 0786/1291] chore(linear_algebra/dimension): fix lemma names (#18747) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This fixes some lemma names which use `rank` but are about `finrank`: * `rank_sup_add_rank_inf_eq` → `submodule.rank_sup_add_rank_inf_eq` * `rank_add_le_rank_add_rank` → `submodule.rank_add_le_rank_add_rank` * `submodule.rank_sup_add_rank_inf_eq` → `submodule.finrank_sup_add_finrank_inf_eq` * `submodule.rank_add_le_rank_add_rank` → `submodule.finrank_add_le_finrank_add_finrank` --- archive/sensitivity.lean | 2 +- src/analysis/inner_product_space/projection.lean | 2 +- .../affine_space/finite_dimensional.lean | 2 +- src/linear_algebra/bilinear_form.lean | 2 +- src/linear_algebra/dimension.lean | 10 +++++----- src/linear_algebra/finite_dimensional.lean | 11 ++++++----- 6 files changed, 15 insertions(+), 14 deletions(-) diff --git a/archive/sensitivity.lean b/archive/sensitivity.lean index 6b3b3e9f6804f..c2c46c393fd79 100644 --- a/archive/sensitivity.lean +++ b/archive/sensitivity.lean @@ -364,7 +364,7 @@ begin { convert ← rank_submodule_le (W ⊔ img), apply dim_V }, have dim_add : dim (W ⊔ img) + dim (W ⊓ img) = dim W + 2^m, - { convert ← rank_sup_add_rank_inf_eq W img, + { convert ← submodule.rank_sup_add_rank_inf_eq W img, rw ← rank_eq_of_injective (g m) g_injective, apply dim_V }, have dimW : dim W = card H, diff --git a/src/analysis/inner_product_space/projection.lean b/src/analysis/inner_product_space/projection.lean index c23076664697c..80d319de5fad2 100644 --- a/src/analysis/inner_product_space/projection.lean +++ b/src/analysis/inner_product_space/projection.lean @@ -1047,7 +1047,7 @@ lemma submodule.finrank_add_inf_finrank_orthogonal {K₁ K₂ : submodule 𝕜 E begin haveI := submodule.finite_dimensional_of_le h, haveI := proper_is_R_or_C 𝕜 K₁, - have hd := submodule.rank_sup_add_rank_inf_eq K₁ (K₁ᗮ ⊓ K₂), + have hd := submodule.finrank_sup_add_finrank_inf_eq K₁ (K₁ᗮ ⊓ K₂), rw [←inf_assoc, (submodule.orthogonal_disjoint K₁).eq_bot, bot_inf_eq, finrank_bot, submodule.sup_orthogonal_inf_of_complete_space h] at hd, rw add_zero at hd, diff --git a/src/linear_algebra/affine_space/finite_dimensional.lean b/src/linear_algebra/affine_space/finite_dimensional.lean index a596e275cdac0..4b009c060e9c8 100644 --- a/src/linear_algebra/affine_space/finite_dimensional.lean +++ b/src/linear_algebra/affine_space/finite_dimensional.lean @@ -681,7 +681,7 @@ begin convert rfl; simp }, { rw [affine_span_coe, direction_affine_span_insert hp₀, add_comm], - refine (submodule.rank_add_le_rank_add_rank _ _).trans (add_le_add_right _ _), + refine (submodule.finrank_add_le_finrank_add_finrank _ _).trans (add_le_add_right _ _), refine finrank_le_one ⟨p -ᵥ p₀, submodule.mem_span_singleton_self _⟩ (λ v, _), have h := v.property, rw submodule.mem_span_singleton at h, diff --git a/src/linear_algebra/bilinear_form.lean b/src/linear_algebra/bilinear_form.lean index aed04531ab05f..68e85e21118e8 100644 --- a/src/linear_algebra/bilinear_form.lean +++ b/src/linear_algebra/bilinear_form.lean @@ -1204,7 +1204,7 @@ begin exact hx₂ n hn }, refine is_compl.of_eq this (eq_top_of_finrank_eq $ (submodule.finrank_le _).antisymm _), conv_rhs { rw ← add_zero (finrank K _) }, - rw [← finrank_bot K V, ← this, submodule.rank_sup_add_rank_inf_eq, + rw [← finrank_bot K V, ← this, submodule.finrank_sup_add_finrank_inf_eq, finrank_add_finrank_orthogonal b₁], exact le_self_add, end diff --git a/src/linear_algebra/dimension.lean b/src/linear_algebra/dimension.lean index 4461772fbe840..9dcc3acba25db 100644 --- a/src/linear_algebra/dimension.lean +++ b/src/linear_algebra/dimension.lean @@ -1113,7 +1113,7 @@ variables [add_comm_group V₂] [module K V₂] variables [add_comm_group V₃] [module K V₃] open linear_map -/-- This is mostly an auxiliary lemma for `rank_sup_add_rank_inf_eq`. -/ +/-- This is mostly an auxiliary lemma for `submodule.rank_sup_add_rank_inf_eq`. -/ lemma rank_add_rank_split (db : V₂ →ₗ[K] V) (eb : V₃ →ₗ[K] V) (cd : V₁ →ₗ[K] V₂) (ce : V₁ →ₗ[K] V₃) (hde : ⊤ ≤ db.range ⊔ eb.range) @@ -1146,7 +1146,7 @@ begin rw [h₂, _root_.neg_neg] } end -lemma rank_sup_add_rank_inf_eq (s t : submodule K V) : +lemma submodule.rank_sup_add_rank_inf_eq (s t : submodule K V) : module.rank K (s ⊔ t : submodule K V) + module.rank K (s ⊓ t : submodule K V) = module.rank K s + module.rank K t := rank_add_rank_split @@ -1165,9 +1165,9 @@ rank_add_rank_split exact ⟨⟨b₁, hb₁, hb₂⟩, rfl, rfl⟩ end -lemma rank_add_le_rank_add_rank (s t : submodule K V) : +lemma submodule.rank_add_le_rank_add_rank (s t : submodule K V) : module.rank K (s ⊔ t : submodule K V) ≤ module.rank K s + module.rank K t := -by { rw [← rank_sup_add_rank_inf_eq], exact self_le_add_right _ _ } +by { rw [← submodule.rank_sup_add_rank_inf_eq], exact self_le_add_right _ _ } end @@ -1360,7 +1360,7 @@ calc rank (f + g) ≤ module.rank K (f.range ⊔ g.range : submodule K V') : assume x, show f x + g x ∈ (f.range ⊔ g.range : submodule K V'), from mem_sup.2 ⟨_, ⟨x, rfl⟩, _, ⟨x, rfl⟩, rfl⟩) end - ... ≤ rank f + rank g : rank_add_le_rank_add_rank _ _ + ... ≤ rank f + rank g : submodule.rank_add_le_rank_add_rank _ _ lemma rank_finset_sum_le {η} (s : finset η) (f : η → V →ₗ[K] V') : rank (∑ d in s, f d) ≤ ∑ d in s, rank (f d) := diff --git a/src/linear_algebra/finite_dimensional.lean b/src/linear_algebra/finite_dimensional.lean index b2fc8bb9ecb9d..617ea765ce6d6 100644 --- a/src/linear_algebra/finite_dimensional.lean +++ b/src/linear_algebra/finite_dimensional.lean @@ -727,7 +727,7 @@ begin end /-- The sum of the dimensions of s + t and s ∩ t is the sum of the dimensions of s and t -/ -theorem rank_sup_add_rank_inf_eq (s t : submodule K V) +theorem finrank_sup_add_finrank_inf_eq (s t : submodule K V) [finite_dimensional K s] [finite_dimensional K t] : finrank K ↥(s ⊔ t) + finrank K ↥(s ⊓ t) = finrank K ↥s + finrank K ↥t := begin @@ -738,10 +738,10 @@ begin exact key end -lemma rank_add_le_rank_add_rank (s t : submodule K V) +lemma finrank_add_le_finrank_add_finrank (s t : submodule K V) [finite_dimensional K s] [finite_dimensional K t] : finrank K (s ⊔ t : submodule K V) ≤ finrank K s + finrank K t := -by { rw [← rank_sup_add_rank_inf_eq], exact self_le_add_right _ _ } +by { rw [← finrank_sup_add_finrank_inf_eq], exact self_le_add_right _ _ } lemma eq_top_of_disjoint [finite_dimensional K V] (s t : submodule K V) (hdim : finrank K s + finrank K t = finrank K V) @@ -752,7 +752,7 @@ begin rw [hdisjoint, finrank_bot] }, apply eq_top_of_finrank_eq, rw ←hdim, - convert s.rank_sup_add_rank_inf_eq t, + convert s.finrank_sup_add_finrank_inf_eq t, rw h_finrank_inf, refl, end @@ -1094,7 +1094,8 @@ lemma finrank_add_eq_of_is_compl [finite_dimensional K V] {U W : submodule K V} (h : is_compl U W) : finrank K U + finrank K W = finrank K V := begin - rw [← rank_sup_add_rank_inf_eq, h.codisjoint.eq_top, h.disjoint.eq_bot, finrank_bot, add_zero], + rw [← finrank_sup_add_finrank_inf_eq, h.codisjoint.eq_top, h.disjoint.eq_bot, finrank_bot, + add_zero], exact finrank_top end From e08a42b2dd544cf11eba72e5fc7bf199d4349925 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 6 Apr 2023 14:44:40 +0000 Subject: [PATCH 0787/1291] =?UTF-8?q?chore(set=5Ftheory/cardinal/basic):?= =?UTF-8?q?=20missing=20lemmas=20about=20`lift=20c=20<=20=E2=84=B5?= =?UTF-8?q?=E2=82=80`=20and=20`=E2=84=B5=E2=82=80=20<=20lift=20c`=20(#1874?= =?UTF-8?q?6)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We already had the `le` versions `lift_le_aleph0` and `aleph0_le_lift`, this adds the `lt` ones `lift_lt_aleph0` and `aleph0_lt_lift`. This turns out to be useful for proving some results about `finrank`, as well as golfing some existing proofs. Since they're trivial, this adds the same lemmas about `continuum` too. --- src/linear_algebra/dimension.lean | 6 ++---- src/set_theory/cardinal/basic.lean | 14 ++++++++++---- src/set_theory/cardinal/continuum.lean | 12 ++++++++++++ src/set_theory/ordinal/arithmetic.lean | 2 +- 4 files changed, 25 insertions(+), 9 deletions(-) diff --git a/src/linear_algebra/dimension.lean b/src/linear_algebra/dimension.lean index 9dcc3acba25db..dcbc43c7acc7f 100644 --- a/src/linear_algebra/dimension.lean +++ b/src/linear_algebra/dimension.lean @@ -867,10 +867,8 @@ by simpa using v.mk_eq_rank lemma basis.nonempty_fintype_index_of_rank_lt_aleph_0 {ι : Type*} (b : basis ι R M) (h : module.rank R M < ℵ₀) : nonempty (fintype ι) := -by rwa [← cardinal.lift_lt, ← b.mk_eq_rank, - -- ensure `aleph_0` has the correct universe - cardinal.lift_aleph_0, ← cardinal.lift_aleph_0.{u_1 v}, - cardinal.lift_lt, cardinal.lt_aleph_0_iff_fintype] at h +by rwa [← cardinal.lift_lt, ← b.mk_eq_rank, cardinal.lift_aleph_0, cardinal.lift_lt_aleph_0, + cardinal.lt_aleph_0_iff_fintype] at h /-- If a module has a finite dimension, all bases are indexed by a finite type. -/ noncomputable def basis.fintype_index_of_rank_lt_aleph_0 {ι : Type*} diff --git a/src/set_theory/cardinal/basic.lean b/src/set_theory/cardinal/basic.lean index 8c1ac0bae9315..c7d1d9905ea81 100644 --- a/src/set_theory/cardinal/basic.lean +++ b/src/set_theory/cardinal/basic.lean @@ -891,6 +891,12 @@ by rw [←lift_aleph_0, lift_le] @[simp] theorem lift_le_aleph_0 {c : cardinal.{u}} : lift.{v} c ≤ ℵ₀ ↔ c ≤ ℵ₀ := by rw [←lift_aleph_0, lift_le] +@[simp] theorem aleph_0_lt_lift {c : cardinal.{u}} : ℵ₀ < lift.{v} c ↔ ℵ₀ < c := +by rw [←lift_aleph_0, lift_lt] + +@[simp] theorem lift_lt_aleph_0 {c : cardinal.{u}} : lift.{v} c < ℵ₀ ↔ c < ℵ₀ := +by rw [←lift_aleph_0, lift_lt] + /-! ### Properties about the cast from `ℕ` -/ @[simp] theorem mk_fin (n : ℕ) : #(fin n) = n := by simp @@ -1245,9 +1251,9 @@ begin apply nat_cast_injective, cases lt_or_ge c ℵ₀ with hc hc, { rw [cast_to_nat_of_lt_aleph_0, ←lift_nat_cast, cast_to_nat_of_lt_aleph_0 hc], - rwa [←lift_aleph_0, lift_lt] }, + rwa [lift_lt_aleph_0] }, { rw [cast_to_nat_of_aleph_0_le, ←lift_nat_cast, cast_to_nat_of_aleph_0_le hc, lift_zero], - rwa [←lift_aleph_0, lift_le] }, + rwa [aleph_0_le_lift] }, end lemma to_nat_congr {β : Type v} (e : α ≃ β) : (#α).to_nat = (#β).to_nat := @@ -1285,8 +1291,8 @@ map_prod to_nat_hom _ _ (ha : a < ℵ₀) (hb : b < ℵ₀) : ((lift.{v u} a) + (lift.{u v} b)).to_nat = a.to_nat + b.to_nat := begin apply cardinal.nat_cast_injective, - replace ha : (lift.{v u} a) < ℵ₀ := by { rw ←lift_aleph_0, exact lift_lt.2 ha }, - replace hb : (lift.{u v} b) < ℵ₀ := by { rw ←lift_aleph_0, exact lift_lt.2 hb }, + replace ha : (lift.{v u} a) < ℵ₀ := by rwa lift_lt_aleph_0, + replace hb : (lift.{u v} b) < ℵ₀ := by rwa lift_lt_aleph_0, rw [nat.cast_add, ←to_nat_lift.{v u} a, ←to_nat_lift.{u v} b, cast_to_nat_of_lt_aleph_0 ha, cast_to_nat_of_lt_aleph_0 hb, cast_to_nat_of_lt_aleph_0 (add_lt_aleph_0 ha hb)] end diff --git a/src/set_theory/cardinal/continuum.lean b/src/set_theory/cardinal/continuum.lean index 8c6fe6fedf417..370883c8c72e3 100644 --- a/src/set_theory/cardinal/continuum.lean +++ b/src/set_theory/cardinal/continuum.lean @@ -39,6 +39,18 @@ by rw [←two_power_aleph_0, lift_two_power, lift_aleph_0, two_power_aleph_0] ### Inequalities -/ +@[simp] lemma continuum_le_lift {c : cardinal.{u}} : 𝔠 ≤ lift.{v} c ↔ 𝔠 ≤ c := +by rw [←lift_continuum, lift_le] + +@[simp] lemma lift_le_continuum {c : cardinal.{u}} : lift.{v} c ≤ 𝔠 ↔ c ≤ 𝔠 := +by rw [←lift_continuum, lift_le] + +@[simp] lemma continuum_lt_lift {c : cardinal.{u}} : 𝔠 < lift.{v} c ↔ 𝔠 < c := +by rw [←lift_continuum, lift_lt] + +@[simp] lemma lift_lt_continuum {c : cardinal.{u}} : lift.{v} c < 𝔠 ↔ c < 𝔠 := +by rw [←lift_continuum, lift_lt] + lemma aleph_0_lt_continuum : ℵ₀ < 𝔠 := cantor ℵ₀ lemma aleph_0_le_continuum : ℵ₀ ≤ 𝔠 := aleph_0_lt_continuum.le diff --git a/src/set_theory/ordinal/arithmetic.lean b/src/set_theory/ordinal/arithmetic.lean index 67d8803ccf422..a860ab829134a 100644 --- a/src/set_theory/ordinal/arithmetic.lean +++ b/src/set_theory/ordinal/arithmetic.lean @@ -1831,7 +1831,7 @@ open ordinal le_antisymm (ord_le.2 $ le_rfl) $ le_of_forall_lt $ λ o h, begin rcases ordinal.lt_lift_iff.1 h with ⟨o, rfl, h'⟩, - rw [lt_ord, ←lift_card, ←lift_aleph_0.{0 u}, lift_lt, ←typein_enum (<) h'], + rw [lt_ord, ←lift_card, lift_lt_aleph_0, ←typein_enum (<) h'], exact lt_aleph_0_iff_fintype.2 ⟨set.fintype_lt_nat _⟩ end From 5ec62c8106221a3f9160e4e4fcc3eed79fe213e9 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 6 Apr 2023 17:41:52 +0000 Subject: [PATCH 0788/1291] feat(linear_algebra/finrank): generalize finrank to ring in more places (#18716) This replaces `[division_ring K]` with the first of the following which compiles: * `[ring K]` * `[ring K] [nontrivial K]` * `[ring K] [strong_rank_condition K]` (implies the previous one via the non-instance `nontrivial_of_invariant_basis_number`) --- .../normed_space/finite_dimension.lean | 4 +- src/category_theory/preadditive/schur.lean | 5 +- src/linear_algebra/eigenspace.lean | 2 +- src/linear_algebra/finite_dimensional.lean | 2 +- src/linear_algebra/finrank.lean | 128 +++++++++++------- 5 files changed, 88 insertions(+), 53 deletions(-) diff --git a/src/analysis/normed_space/finite_dimension.lean b/src/analysis/normed_space/finite_dimension.lean index 2aeddb53260b9..0252f919cb9f8 100644 --- a/src/analysis/normed_space/finite_dimension.lean +++ b/src/analysis/normed_space/finite_dimension.lean @@ -398,7 +398,7 @@ explicitly when needed. -/ variables (𝕜 E) lemma finite_dimensional.complete [finite_dimensional 𝕜 E] : complete_space E := begin - set e := continuous_linear_equiv.of_finrank_eq (@finrank_fin_fun 𝕜 _ (finrank 𝕜 E)).symm, + set e := continuous_linear_equiv.of_finrank_eq (@finrank_fin_fun 𝕜 _ _ (finrank 𝕜 E)).symm, have : uniform_embedding e.to_linear_equiv.to_equiv.symm := e.symm.uniform_embedding, exact (complete_space_congr this).1 (by apply_instance) end @@ -626,7 +626,7 @@ properness of `𝕜`, and the search for `𝕜` as an unknown metavariable. Decl explicitly when needed. -/ lemma finite_dimensional.proper [finite_dimensional 𝕜 E] : proper_space E := begin - set e := continuous_linear_equiv.of_finrank_eq (@finrank_fin_fun 𝕜 _ (finrank 𝕜 E)).symm, + set e := continuous_linear_equiv.of_finrank_eq (@finrank_fin_fun 𝕜 _ _ (finrank 𝕜 E)).symm, exact e.symm.antilipschitz.proper_space e.symm.continuous e.symm.surjective end diff --git a/src/category_theory/preadditive/schur.lean b/src/category_theory/preadditive/schur.lean index 2faa96f55b395..533e1234c484f 100644 --- a/src/category_theory/preadditive/schur.lean +++ b/src/category_theory/preadditive/schur.lean @@ -123,8 +123,7 @@ lemma finrank_endomorphism_eq_one finrank 𝕜 (X ⟶ X) = 1 := begin have id_nonzero := (is_iso_iff_nonzero (𝟙 X)).mp (by apply_instance), - apply finrank_eq_one (𝟙 X), - { exact id_nonzero, }, + refine finrank_eq_one (𝟙 X) id_nonzero _, { intro f, haveI : nontrivial (End X) := nontrivial_of_ne _ _ id_nonzero, obtain ⟨c, nu⟩ := @spectrum.nonempty_of_is_alg_closed_of_finite_dimensional 𝕜 (End X) _ _ _ _ _ @@ -184,7 +183,7 @@ begin exact zero_le_one }, { obtain ⟨f, nz⟩ := (nontrivial_iff_exists_ne 0).mp h, haveI fi := (is_iso_iff_nonzero f).mpr nz, - apply finrank_le_one f, + refine finrank_le_one f _, intro g, obtain ⟨c, w⟩ := endomorphism_simple_eq_smul_id 𝕜 (g ≫ inv f), exact ⟨c, by simpa using w =≫ f⟩, }, diff --git a/src/linear_algebra/eigenspace.lean b/src/linear_algebra/eigenspace.lean index 59b617cdef3d6..f1551a369b413 100644 --- a/src/linear_algebra/eigenspace.lean +++ b/src/linear_algebra/eigenspace.lean @@ -541,7 +541,7 @@ begin cases n, -- If the vector space is 0-dimensional, the result is trivial. { rw ←top_le_iff, - simp only [finrank_eq_zero.1 (eq.trans finrank_top h_dim), bot_le] }, + simp only [finrank_eq_zero.1 (eq.trans (finrank_top _ _) h_dim), bot_le] }, -- Otherwise the vector space is nontrivial. { haveI : nontrivial V := finrank_pos_iff.1 (by { rw h_dim, apply nat.zero_lt_succ }), -- Hence, `f` has an eigenvalue `μ₀`. diff --git a/src/linear_algebra/finite_dimensional.lean b/src/linear_algebra/finite_dimensional.lean index 617ea765ce6d6..8caa119dccdbe 100644 --- a/src/linear_algebra/finite_dimensional.lean +++ b/src/linear_algebra/finite_dimensional.lean @@ -1096,7 +1096,7 @@ lemma finrank_add_eq_of_is_compl begin rw [← finrank_sup_add_finrank_inf_eq, h.codisjoint.eq_top, h.disjoint.eq_bot, finrank_bot, add_zero], - exact finrank_top + exact finrank_top _ _ end end division_ring diff --git a/src/linear_algebra/finrank.lean b/src/linear_algebra/finrank.lean index 5aaafab257c8c..baf0f4822cd01 100644 --- a/src/linear_algebra/finrank.lean +++ b/src/linear_algebra/finrank.lean @@ -40,9 +40,8 @@ namespace finite_dimensional open is_noetherian -section division_ring - -variables [division_ring K] [add_comm_group V] [module K V] +section ring +variables [ring K] [add_comm_group V] [module K V] {V₂ : Type v'} [add_comm_group V₂] [module K V₂] /-- The rank of a module as a natural number. @@ -86,26 +85,18 @@ begin exact n.zero_le }, end -/-- 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) : - finrank K V = fintype.card ι := -finrank_eq_of_rank_eq (rank_eq_card_basis h) - -/-- If a vector space has a finite basis, then its dimension is equal to the cardinality of the -basis. This lemma uses a `finset` instead of indexed types. -/ -lemma finrank_eq_card_finset_basis {ι : Type w} {b : finset ι} - (h : basis.{w} b K V) : - finrank K V = finset.card b := -by rw [finrank_eq_card_basis h, fintype.card_coe] +section +variables [nontrivial K] [no_zero_smul_divisors K V] /-- A finite dimensional space is nontrivial if it has positive `finrank`. -/ -lemma nontrivial_of_finrank_pos (h : 0 < finrank K V) : nontrivial V := +lemma nontrivial_of_finrank_pos (h : 0 < finrank K V) : + nontrivial V := rank_pos_iff_nontrivial.mp (rank_lt_of_finrank_lt h) /-- A finite dimensional space is nontrivial if it has `finrank` equal to the successor of a natural number. -/ -lemma nontrivial_of_finrank_eq_succ {n : ℕ} (hn : finrank K V = n.succ) : nontrivial V := +lemma nontrivial_of_finrank_eq_succ {n : ℕ} + (hn : finrank K V = n.succ) : nontrivial V := nontrivial_of_finrank_pos (by rw hn; exact n.succ_pos) /-- A (finite dimensional) space that is a subsingleton has zero `finrank`. -/ @@ -117,12 +108,28 @@ begin exact hxy (subsingleton.elim _ _) end -lemma basis.subset_extend {s : set V} (hs : linear_independent K (coe : s → V)) : - s ⊆ hs.extend (set.subset_univ _) := -hs.subset_extend _ +end + +section +variables [strong_rank_condition K] + +/-- If a vector space (or module) has a finite basis, then its dimension (or rank) is equal to the +cardinality of the basis. -/ +lemma finrank_eq_card_basis {ι : Type w} [fintype ι] (h : basis ι K V) : + finrank K V = fintype.card ι := +finrank_eq_of_rank_eq (rank_eq_card_basis h) + +/-- If a vector space (or module) has a finite basis, then its dimension (or rank) is equal to the +cardinality of the basis. This lemma uses a `finset` instead of indexed types. -/ +lemma finrank_eq_card_finset_basis {ι : Type w} {b : finset ι} + (h : basis.{w} b K V) : + finrank K V = finset.card b := +by rw [finrank_eq_card_basis h, fintype.card_coe] variable (K) -/-- A division_ring is one-dimensional as a vector space over itself. -/ + +/-- A ring satisfying `strong_rank_condition` (such as a `division_ring`) is one-dimensional as a +module over itself. -/ @[simp] lemma finrank_self : finrank K K = 1 := finrank_eq_of_rank_eq (by simp) @@ -135,6 +142,19 @@ finrank_eq_of_rank_eq rank_fun' @[simp] lemma finrank_fin_fun {n : ℕ} : finrank K (fin n → K) = n := by simp +end + +end ring + +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 basis.subset_extend {s : set V} (hs : linear_independent K (coe : s → V)) : + s ⊆ hs.extend (set.subset_univ _) := +hs.subset_extend _ + end division_ring end finite_dimensional @@ -143,14 +163,16 @@ variables {K V} section zero_rank -variables [division_ring K] [add_comm_group V] [module K V] +variables [ring K] [strong_rank_condition K] [add_comm_group V] [module K V] [module.free K V] open finite_dimensional lemma finrank_eq_zero_of_basis_imp_not_finite (h : ∀ s : set V, basis.{v} (s : set V) K V → ¬ s.finite) : finrank K V = 0 := -dif_neg (λ rank_lt, h _ (basis.of_vector_space K V) - ((basis.of_vector_space K V).finite_index_of_rank_lt_aleph_0 rank_lt)) +begin + obtain ⟨_, ⟨b⟩⟩ := (module.free_iff_set K V).mp ‹_›, + exact dif_neg (λ rank_lt, h _ b (b.finite_index_of_rank_lt_aleph_0 rank_lt)) +end lemma finrank_eq_zero_of_basis_imp_false (h : ∀ s : finset V, basis.{v} (s : set V) K V → false) : finrank K V = 0 := @@ -168,17 +190,12 @@ lemma finrank_eq_zero_of_not_exists_basis_finset (h : ¬ ∃ (s : finset V), nonempty (basis s K V)) : finrank K V = 0 := finrank_eq_zero_of_basis_imp_false (λ s b, h ⟨s, ⟨b⟩⟩) -variables (K V) - -@[simp] lemma finrank_bot : finrank K (⊥ : submodule K V) = 0 := -finrank_eq_of_rank_eq (rank_bot _ _) - end zero_rank namespace linear_equiv open finite_dimensional -variables [division_ring K] [add_comm_group V] [module K V] +variables [ring K] [add_comm_group V] [module K V] {V₂ : Type v'} [add_comm_group V₂] [module K V₂] variables {R M M₂ : Type*} [ring R] [add_comm_group M] [add_comm_group M₂] @@ -198,8 +215,8 @@ end linear_equiv namespace linear_map open finite_dimensional -section division_ring -variables [division_ring K] [add_comm_group V] [module K V] +section ring +variables [ring K] [add_comm_group V] [module K V] {V₂ : Type v'} [add_comm_group V₂] [module K V₂] /-- The dimensions of the domain and range of an injective linear map are equal. -/ @@ -207,14 +224,19 @@ 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 +end ring end linear_map open module finite_dimensional section -variables [division_ring K] [add_comm_group V] [module K V] +variables [ring K] [add_comm_group V] [module K V] + +variables (K V) + +@[simp] lemma finrank_bot [nontrivial K] : finrank K (⊥ : submodule K V) = 0 := +finrank_eq_of_rank_eq (rank_bot _ _) @[simp] theorem finrank_top : finrank K (⊤ : submodule K V) = finrank K V := @@ -224,8 +246,8 @@ end namespace submodule -section division_ring -variables [division_ring K] [add_comm_group V] [module K V] +section ring +variables [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} @@ -235,11 +257,11 @@ lt_of_le_of_ne le (λ h, ne_of_lt lt (by rw h)) lemma lt_top_of_finrank_lt_finrank {s : submodule K V} (lt : finrank K s < finrank K V) : s < ⊤ := begin - rw ← @finrank_top K V at lt, + rw ← finrank_top K V at lt, exact lt_of_le_of_finrank_lt_finrank le_top lt end -end division_ring +end ring end submodule @@ -430,13 +452,16 @@ 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] +variables [ring K] [add_comm_group V] [module K V] +variables [no_zero_smul_divisors K V] [strong_rank_condition K] /-- 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) : +lemma finrank_eq_one + (v : V) (n : v ≠ 0) (h : ∀ w : V, ∃ c : K, c • v = w) : finrank K V = 1 := begin + haveI := nontrivial_of_invariant_basis_number K, obtain ⟨b⟩ := (basis.basis_singleton_iff punit).mpr ⟨v, n, h⟩, rw [finrank_eq_card_basis b, fintype.card_punit] end @@ -447,6 +472,7 @@ If every vector is a multiple of some `v : V`, then `V` has dimension at most on lemma finrank_le_one (v : V) (h : ∀ w : V, ∃ c : K, c • v = w) : finrank K V ≤ 1 := begin + haveI := nontrivial_of_invariant_basis_number K, rcases eq_or_ne v 0 with rfl | hn, { haveI := subsingleton_of_forall_eq (0 : V) (λ w, by { obtain ⟨c, rfl⟩ := h w, simp }), rw finrank_zero_of_subsingleton, @@ -458,12 +484,8 @@ end finrank_eq_one section subalgebra_rank open module -variables {F E : Type*} [field F] [ring E] [algebra F E] -@[simp] lemma subalgebra.rank_bot [nontrivial E] : module.rank F (⊥ : subalgebra F E) = 1 := -((subalgebra.to_submodule_equiv (⊥ : subalgebra F E)).symm.trans $ - linear_equiv.of_eq _ _ algebra.to_submodule_bot).rank_eq.trans $ - by { rw rank_span_set, exacts [mk_singleton _, linear_independent_singleton one_ne_zero] } +variables {F E : Type*} [comm_ring F] [ring E] [algebra F E] @[simp] lemma subalgebra.rank_to_submodule (S : subalgebra F E) : module.rank F S.to_submodule = module.rank F S := rfl @@ -482,8 +504,22 @@ by { rw ← algebra.top_to_submodule, refl } lemma subalgebra.rank_top : module.rank F (⊤ : subalgebra F E) = module.rank F E := by { rw subalgebra_top_rank_eq_submodule_top_rank, exact rank_top F E } +section +variables [strong_rank_condition F] [no_zero_smul_divisors F E] [nontrivial E] + +@[simp] lemma subalgebra.rank_bot : + module.rank F (⊥ : subalgebra F E) = 1 := +((subalgebra.to_submodule_equiv (⊥ : subalgebra F E)).symm.trans $ + linear_equiv.of_eq _ _ algebra.to_submodule_bot).rank_eq.trans $ begin + letI := module.nontrivial F E, + rw rank_span_set, + exacts [mk_singleton _, linear_independent_singleton one_ne_zero] + end + @[simp] -lemma subalgebra.finrank_bot [nontrivial E] : finrank F (⊥ : subalgebra F E) = 1 := +lemma subalgebra.finrank_bot : finrank F (⊥ : subalgebra F E) = 1 := finrank_eq_of_rank_eq (by simp) +end + end subalgebra_rank From 9dba31df156d9d65b9d78db449542ca73d147c68 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 7 Apr 2023 10:00:47 +0000 Subject: [PATCH 0789/1291] =?UTF-8?q?chore(set=5Ftheory/cardinal/basic):?= =?UTF-8?q?=20missing=20lemmas=20about=20`<`/`=E2=89=A4`=20and=20`lift=20c?= =?UTF-8?q?`/`=E2=86=91n`=20(#18752)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- src/set_theory/cardinal/basic.lean | 14 ++++++++++++++ 1 file changed, 14 insertions(+) diff --git a/src/set_theory/cardinal/basic.lean b/src/set_theory/cardinal/basic.lean index c7d1d9905ea81..f0028ca8d6974 100644 --- a/src/set_theory/cardinal/basic.lean +++ b/src/set_theory/cardinal/basic.lean @@ -911,6 +911,20 @@ lift_injective.eq_iff' (lift_nat_cast n) (n : cardinal) = lift.{v} a ↔ (n : cardinal) = a := by rw [←lift_nat_cast.{v} n, lift_inj] +@[simp] lemma lift_le_nat_iff {a : cardinal.{u}} {n : ℕ} : lift.{v} a ≤ n ↔ a ≤ n := +by simp only [←lift_nat_cast, lift_le] + +@[simp] lemma nat_le_lift_iff {n : ℕ} {a : cardinal.{u}} : + (n : cardinal) ≤ lift.{v} a ↔ (n : cardinal) ≤ a := +by simp only [←lift_nat_cast, lift_le] + +@[simp] lemma lift_lt_nat_iff {a : cardinal.{u}} {n : ℕ} : lift.{v} a < n ↔ a < n := +by simp only [←lift_nat_cast, lift_lt] + +@[simp] lemma nat_lt_lift_iff {n : ℕ} {a : cardinal.{u}} : + (n : cardinal) < lift.{v} a ↔ (n : cardinal) < a := +by simp only [←lift_nat_cast, lift_lt] + theorem lift_mk_fin (n : ℕ) : lift (#(fin n)) = n := by simp lemma mk_coe_finset {α : Type u} {s : finset α} : #s = ↑(finset.card s) := by simp From dc65937e7a0fff4677f0f5a7086f42da70a69575 Mon Sep 17 00:00:00 2001 From: Michael Stoll Date: Fri, 7 Apr 2023 15:31:22 +0000 Subject: [PATCH 0790/1291] feat(number_theory/pell): add API for solutions (#18626) This continues the devlopment of the theory of Pell's equation. We add some API lemmas for solutions, which will be needed for defining of the fundamental solution and proving its properties. See [this Zulip thread](https://leanprover.zulipchat.com/#narrow/stream/217875-Is-there-code-for-X.3F/topic/Proving.20Pell's.20equation.20is.20solvable/near/343270338). --- src/number_theory/pell.lean | 168 ++++++++++++++++++++++++++++++++++-- 1 file changed, 160 insertions(+), 8 deletions(-) diff --git a/src/number_theory/pell.lean b/src/number_theory/pell.lean index f6fb65d95dae1..126d65f2a2b8e 100644 --- a/src/number_theory/pell.lean +++ b/src/number_theory/pell.lean @@ -12,16 +12,27 @@ import number_theory.zsqrtd.basic /-! # Pell's Equation -We prove the following +*Pell's Equation* is the equation $x^2 - d y^2 = 1$, where $d$ is a positive integer +that is not a square, and one is interested in solutions in integers $x$ and $y$. + +In this file, we aim at providing all of the essential theory of Pell's Equation for general $d$ +(as opposed to the contents of `number_theory.pell_matiyasevic`, which is specific to the case +$d = a^2 - 1$ for some $a > 1$). + +We begin by defining a type `pell.solution₁ d` for solutions of the equation, +show that it has a natural structure as an abelian group, and prove some basic +properties. + +We then prove the following **Theorem.** Let $d$ be a positive integer that is not a square. Then the equation $x^2 - d y^2 = 1$ has a nontrivial (i.e., with $y \ne 0$) solution in integers. -See `pell.exists_of_not_is_square`. +See `pell.exists_of_not_is_square` and `pell.exists_nontrivial_of_not_is_square`. -This is the beginning of a development that aims at providing all of the essential theory -of Pell's Equation for general $d$ (as opposed to the contents of `number_theory.pell_matiyasevic`, -which is specific to the case $d = a^2 - 1$ for some $a > 1$). +The next step (TODO) will be to define the *fundamental solution* as the solution +with smallest $x$ among all solutions satisfying $x > 1$ and $y > 0$ and to show +that every solution is a power of the fundamental solution up to a (common) sign. ## References @@ -144,6 +155,133 @@ lemma x_neg (a : solution₁ d) : (-a).x = -a.x := rfl @[simp] lemma y_neg (a : solution₁ d) : (-a).y = -a.y := rfl +/-- When `d` is negative, then `x` or `y` must be zero in a solution. -/ +lemma eq_zero_of_d_neg (h₀ : d < 0) (a : solution₁ d) : a.x = 0 ∨ a.y = 0 := +begin + have h := a.prop, + contrapose! h, + have h1 := sq_pos_of_ne_zero a.x h.1, + have h2 := sq_pos_of_ne_zero a.y h.2, + nlinarith, +end + +/-- A solution has `x ≠ 0`. -/ +lemma x_ne_zero (h₀ : 0 ≤ d) (a : solution₁ d) : a.x ≠ 0 := +begin + intro hx, + have h : 0 ≤ d * a.y ^ 2 := mul_nonneg h₀ (sq_nonneg _), + rw [a.prop_y, hx, sq, zero_mul, zero_sub] at h, + exact not_le.mpr (neg_one_lt_zero : (-1 : ℤ) < 0) h, +end + +/-- A solution with `x > 1` must have `y ≠ 0`. -/ +lemma y_ne_zero_of_one_lt_x {a : solution₁ d} (ha : 1 < a.x) : a.y ≠ 0 := +begin + intro hy, + have prop := a.prop, + rw [hy, sq (0 : ℤ), zero_mul, mul_zero, sub_zero] at prop, + exact lt_irrefl _ (((one_lt_sq_iff $ zero_le_one.trans ha.le).mpr ha).trans_eq prop), +end + +/-- A solution with `x = 1` is trivial. -/ +lemma eq_one_of_x_eq_one (h₀ : d ≠ 0) {a : solution₁ d} (ha : a.x = 1) : a = 1 := +begin + have prop := a.prop_y, + rw [ha, one_pow, sub_self, mul_eq_zero, or_iff_right h₀, sq_eq_zero_iff] at prop, + exact ext ha prop, +end + +/-- A solution is `1` or `-1` if and only if `y = 0`. -/ +lemma eq_one_or_neg_one_iff_y_eq_zero {a : solution₁ d} : a = 1 ∨ a = -1 ↔ a.y = 0 := +begin + refine ⟨λ H, H.elim (λ h, by simp [h]) (λ h, by simp [h]), λ H, _⟩, + have prop := a.prop, + rw [H, sq (0 : ℤ), mul_zero, mul_zero, sub_zero, sq_eq_one_iff] at prop, + exact prop.imp (λ h, ext h H) (λ h, ext h H), +end + +/-- The set of solutions with `x > 0` is closed under multiplication. -/ +lemma x_mul_pos {a b : solution₁ d} (ha : 0 < a.x) (hb : 0 < b.x) : 0 < (a * b).x := +begin + simp only [x_mul], + refine neg_lt_iff_pos_add'.mp (abs_lt.mp _).1, + rw [← abs_of_pos ha, ← abs_of_pos hb, ← abs_mul, ← sq_lt_sq, mul_pow a.x, a.prop_x, b.prop_x, + ← sub_pos], + ring_nf, + cases le_or_lt 0 d with h h, + { positivity, }, + { rw [(eq_zero_of_d_neg h a).resolve_left ha.ne', (eq_zero_of_d_neg h b).resolve_left hb.ne', + zero_pow two_pos, zero_add, zero_mul, zero_add], + exact one_pos, }, +end + +/-- The set of solutions with `x` and `y` positive is closed under multiplication. -/ +lemma y_mul_pos {a b : solution₁ d} (hax : 0 < a.x) (hay : 0 < a.y) (hbx : 0 < b.x) + (hby : 0 < b.y) : + 0 < (a * b).y := +begin + simp only [y_mul], + positivity, +end + +/-- If `(x, y)` is a solution with `x` positive, then all its powers with natural exponents +have positive `x`. -/ +lemma x_pow_pos {a : solution₁ d} (hax : 0 < a.x) (n : ℕ) : 0 < (a ^ n).x := +begin + induction n with n ih, + { simp only [pow_zero, x_one, zero_lt_one], }, + { rw [pow_succ], + exact x_mul_pos hax ih, } +end + +/-- If `(x, y)` is a solution with `x` and `y` positive, then all its powers with positive +natural exponents have positive `y`. -/ +lemma y_pow_succ_pos {a : solution₁ d} (hax : 0 < a.x) (hay : 0 < a.y) (n : ℕ) : + 0 < (a ^ n.succ).y := +begin + induction n with n ih, + { simp only [hay, pow_one], }, + { rw [pow_succ], + exact y_mul_pos hax hay (x_pow_pos hax _) ih, } +end + +/-- If `(x, y)` is a solution with `x` positive, then all its powers have positive `x`. -/ +lemma x_zpow_pos {a : solution₁ d} (hax : 0 < a.x) (n : ℤ) : 0 < (a ^ n).x := +begin + cases n, + { rw zpow_of_nat, + exact x_pow_pos hax n }, + { rw zpow_neg_succ_of_nat, + exact x_pow_pos hax (n + 1) }, +end + +/-- If `(x, y)` is a solution with `x` and `y` positive, then the `y` component of any power +has the same sign as the exponent. -/ +lemma sign_y_zpow_eq_sign_of_x_pos_of_y_pos {a : solution₁ d} (hax : 0 < a.x) (hay : 0 < a.y) + (n : ℤ) : + (a ^ n).y.sign = n.sign := +begin + rcases n with (_ | _) | _, + { refl }, + { rw zpow_of_nat, + exact int.sign_eq_one_of_pos (y_pow_succ_pos hax hay n) }, + { rw zpow_neg_succ_of_nat, + exact int.sign_eq_neg_one_of_neg (neg_neg_of_pos (y_pow_succ_pos hax hay n)) }, +end + +/-- If `a` is any solution, then one of `a`, `a⁻¹`, `-a`, `-a⁻¹` has +positive `x` and nonnegative `y`. -/ +lemma exists_pos_variant (h₀ : 0 < d) (a : solution₁ d) : + ∃ b : solution₁ d, 0 < b.x ∧ 0 ≤ b.y ∧ a ∈ ({b, b⁻¹, -b, -b⁻¹} : set (solution₁ d)) := +begin + refine (lt_or_gt_of_ne (a.x_ne_zero h₀.le)).elim + ((le_total 0 a.y).elim (λ hy hx, ⟨-a⁻¹, _, _, _⟩) (λ hy hx, ⟨-a, _, _, _⟩)) + ((le_total 0 a.y).elim (λ hy hx, ⟨a, hx, hy, _⟩) (λ hy hx, ⟨a⁻¹, hx, _, _⟩)); + simp only [neg_neg, inv_inv, neg_inv, set.mem_insert_iff, set.mem_singleton_iff, true_or, + eq_self_iff_true, x_neg, x_inv, y_neg, y_inv, neg_pos, neg_nonneg, or_true]; + assumption, +end + end solution₁ section existence @@ -234,16 +372,30 @@ begin simpa [mul_self_pos.mp h₀, sub_eq_add_neg, eq_neg_self_iff] using int.eq_of_mul_eq_one hxy, end +namespace solution₁ + +/-- If `d` is a positive integer that is not a square, then there exists a nontrivial solution +to the Pell equation `x^2 - d*y^2 = 1`. -/ +theorem exists_nontrivial_of_not_is_square (h₀ : 0 < d) (hd : ¬ is_square d) : + ∃ a : solution₁ d, a ≠ 1 ∧ a ≠ -1 := +begin + obtain ⟨x, y, prop, hy⟩ := exists_of_not_is_square h₀ hd, + refine ⟨mk x y prop, λ H, _, λ H, _⟩; apply_fun solution₁.y at H; simpa only [hy] using H, +end + /-- If `d` is a positive integer that is not a square, then there exists a solution to the Pell equation `x^2 - d*y^2 = 1` with `x > 1` and `y > 0`. -/ lemma exists_pos_of_not_is_square (h₀ : 0 < d) (hd : ¬ is_square d) : - ∃ x y : ℤ, x ^ 2 - d * y ^ 2 = 1 ∧ 1 < x ∧ 0 < y := + ∃ a : solution₁ d, 1 < a.x ∧ 0 < a.y := begin obtain ⟨x, y, h, hy⟩ := exists_of_not_is_square h₀ hd, - refine ⟨|x|, |y|, by rwa [sq_abs, sq_abs], _, abs_pos.mpr hy⟩, - rw [← one_lt_sq_iff_one_lt_abs, eq_add_of_sub_eq h, lt_add_iff_pos_right], + refine ⟨mk (|x|) (|y|) (by rwa [sq_abs, sq_abs]), _, abs_pos.mpr hy⟩, + rw [x_mk, ← one_lt_sq_iff_one_lt_abs, eq_add_of_sub_eq h, lt_add_iff_pos_right], exact mul_pos h₀ (sq_pos_of_ne_zero y hy), end +end solution₁ + end existence + end pell From f29120f82f6e24a6f6579896dfa2de6769fec962 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Fri, 7 Apr 2023 22:13:52 +0000 Subject: [PATCH 0791/1291] chore(order/rel_iso/basic): better `namespace` management (#18758) We remove a lot of `_root_` by simply closing and reopening a namespace. --- src/order/rel_iso/basic.lean | 29 ++++++++++++++++------------- 1 file changed, 16 insertions(+), 13 deletions(-) diff --git a/src/order/rel_iso/basic.lean b/src/order/rel_iso/basic.lean index ab68f286cbf01..2c78215fa6407 100644 --- a/src/order/rel_iso/basic.lean +++ b/src/order/rel_iso/basic.lean @@ -306,20 +306,22 @@ protected theorem is_well_founded (f : r ↪r s) [is_well_founded β s] : is_wel protected theorem is_well_order : ∀ (f : r ↪r s) [is_well_order β s], is_well_order α r | f H := by exactI {wf := f.well_founded H.wf, ..f.is_strict_total_order} -instance _root_.subtype.well_founded_lt [has_lt α] [well_founded_lt α] (p : α → Prop) : +end rel_embedding + +instance subtype.well_founded_lt [has_lt α] [well_founded_lt α] (p : α → Prop) : well_founded_lt (subtype p) := (subtype.rel_embedding (<) p).is_well_founded -instance _root_.subtype.well_founded_gt [has_lt α] [well_founded_gt α] (p : α → Prop) : +instance subtype.well_founded_gt [has_lt α] [well_founded_gt α] (p : α → Prop) : well_founded_gt (subtype p) := (subtype.rel_embedding (>) p).is_well_founded /-- `quotient.mk` as a relation homomorphism between the relation and the lift of a relation. -/ -@[simps] def _root_.quotient.mk_rel_hom [setoid α] {r : α → α → Prop} (H) : +@[simps] def quotient.mk_rel_hom [setoid α] {r : α → α → Prop} (H) : r →r quotient.lift₂ r H := ⟨@quotient.mk α _, λ _ _, id⟩ /-- `quotient.out` as a relation embedding between the lift of a relation and the relation. -/ @[simps] -noncomputable def _root_.quotient.out_rel_embedding [setoid α] {r : α → α → Prop} (H) : +noncomputable def quotient.out_rel_embedding [setoid α] {r : α → α → Prop} (H) : quotient.lift₂ r H ↪r r := ⟨embedding.quotient_out α, begin refine λ x y, quotient.induction_on₂ x y (λ a b, _), @@ -329,12 +331,12 @@ end⟩ /-- `quotient.out'` as a relation embedding between the lift of a relation and the relation. -/ @[simps] -noncomputable def _root_.quotient.out'_rel_embedding {s : setoid α} {r : α → α → Prop} (H) : +noncomputable def quotient.out'_rel_embedding {s : setoid α} {r : α → α → Prop} (H) : (λ a b, quotient.lift_on₂' a b r H) ↪r r := { to_fun := quotient.out', ..quotient.out_rel_embedding _ } -@[simp] theorem _root_.acc_lift₂_iff [setoid α] {r : α → α → Prop} {H} {a} : +@[simp] theorem acc_lift₂_iff [setoid α] {r : α → α → Prop} {H} {a} : acc (quotient.lift₂ r H) ⟦a⟧ ↔ acc r a := begin split, @@ -346,12 +348,12 @@ begin exact IH a' h, }, end -@[simp] theorem _root_.acc_lift_on₂'_iff {s : setoid α} {r : α → α → Prop} {H} {a} : +@[simp] theorem acc_lift_on₂'_iff {s : setoid α} {r : α → α → Prop} {H} {a} : acc (λ x y, quotient.lift_on₂' x y r H) (quotient.mk' a : quotient s) ↔ acc r a := acc_lift₂_iff /-- A relation is well founded iff its lift to a quotient is. -/ -theorem _root_.well_founded_lift₂_iff [setoid α] {r : α → α → Prop} {H} : +theorem well_founded_lift₂_iff [setoid α] {r : α → α → Prop} {H} : well_founded (quotient.lift₂ r H) ↔ well_founded r := begin split, @@ -361,15 +363,16 @@ begin exact acc_lift₂_iff.2 (wf.apply a), }, end -alias _root_.well_founded_lift₂_iff ↔ - _root_.well_founded.of_quotient_lift₂ _root_.well_founded.quotient_lift₂ +alias well_founded_lift₂_iff ↔ well_founded.of_quotient_lift₂ well_founded.quotient_lift₂ -@[simp] theorem _root_.well_founded_lift_on₂'_iff {s : setoid α} {r : α → α → Prop} {H} : +@[simp] theorem well_founded_lift_on₂'_iff {s : setoid α} {r : α → α → Prop} {H} : well_founded (λ x y : quotient s, quotient.lift_on₂' x y r H) ↔ well_founded r := well_founded_lift₂_iff -alias _root_.well_founded_lift_on₂'_iff ↔ - _root_.well_founded.of_quotient_lift_on₂' _root_.well_founded.quotient_lift_on₂' +alias well_founded_lift_on₂'_iff ↔ + well_founded.of_quotient_lift_on₂' well_founded.quotient_lift_on₂' + +namespace rel_embedding /-- To define an relation embedding from an antisymmetric relation `r` to a reflexive relation `s` it From 3a2039ad20626aebbece10cddb18ef010f7d912d Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sat, 8 Apr 2023 07:53:36 +0000 Subject: [PATCH 0792/1291] chore(linear_algebra/dimension): remove outdated comment (#18759) This used to be about `vector_space.dim`, but was caught up in the rename that fixed the naming inconsistency and is now nonsense. --- src/linear_algebra/dimension.lean | 4 ---- 1 file changed, 4 deletions(-) diff --git a/src/linear_algebra/dimension.lean b/src/linear_algebra/dimension.lean index dcbc43c7acc7f..55283a585b87e 100644 --- a/src/linear_algebra/dimension.lean +++ b/src/linear_algebra/dimension.lean @@ -64,10 +64,6 @@ For vector spaces (i.e. modules over a field), we have ## Implementation notes -There is a naming discrepancy: most of the theorem names refer to `rank`, -even though the definition is of `module.rank`. -This reflects that `module.rank` was originally called `rank`, and only defined for vector spaces. - Many theorems in this file are not universe-generic when they relate dimensions in different universes. They should be as general as they can be without inserting `lift`s. The types `V`, `V'`, ... all live in different universes, From 2efd2423f8d25fa57cf7a179f5d8652ab4d0df44 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Sat, 8 Apr 2023 11:36:43 +0000 Subject: [PATCH 0793/1291] chore(category_theory): simps should not add hom lemmas (#18742) `@[simps]` should not be used to simplify the `hom` field of a category instance. Very little needs to be changed when removing it. However the problem in https://github.com/leanprover-community/mathlib4/pull/3244 with a proof by `simp` failing seems to be implicated by this problem. If we remove the `@[simps]` generated lemma for `Hom` there, the original proof works (although is extremely slow). Co-authored-by: Scott Morrison --- src/category_theory/category/basic.lean | 6 ++++-- src/category_theory/fin_category.lean | 5 ++--- src/category_theory/groupoid.lean | 5 +++++ src/category_theory/monoidal/braided.lean | 3 +++ src/category_theory/sites/sheaf.lean | 2 +- 5 files changed, 15 insertions(+), 6 deletions(-) diff --git a/src/category_theory/category/basic.lean b/src/category_theory/category/basic.lean index 02f3ac93913f6..6a119b1d4800e 100644 --- a/src/category_theory/category/basic.lean +++ b/src/category_theory/category/basic.lean @@ -85,6 +85,8 @@ extends quiver.{v+1} obj : Type (max u (v+1)) := notation `𝟙` := category_struct.id -- type as \b1 infixr ` ≫ `:80 := category_struct.comp -- type as \gg +initialize_simps_projections category_struct (-to_quiver_hom) + /-- The typeclass `category C` describes morphisms associated to objects of type `C`. The universe levels of the objects and morphisms are unconstrained, and will often need to be @@ -122,8 +124,8 @@ abbreviation small_category (C : Type u) : Type (u+1) := category.{u} C section variables {C : Type u} [category.{v} C] {X Y Z : C} -initialize_simps_projections category (to_category_struct_to_quiver_hom → hom, - to_category_struct_comp → comp, to_category_struct_id → id, -to_category_struct) +initialize_simps_projections category + (to_category_struct_comp → comp, to_category_struct_id → id, -to_category_struct) /-- postcompose an equation between morphisms by another morphism -/ lemma eq_whisker {f g : X ⟶ Y} (w : f = g) (h : Y ⟶ Z) : f ≫ h = g ≫ h := diff --git a/src/category_theory/fin_category.lean b/src/category_theory/fin_category.lean index 0681cdda833d7..a9ba4f285edeb 100644 --- a/src/category_theory/fin_category.lean +++ b/src/category_theory/fin_category.lean @@ -60,15 +60,14 @@ noncomputable def obj_as_type_equiv : obj_as_type α ≌ α := /-- A fin_category `α` is equivalent to a fin_category with in `Type`. -/ @[nolint unused_arguments] abbreviation as_type : Type := fin (fintype.card α) -@[simps hom id comp (lemmas_only)] noncomputable +@[simps id comp (lemmas_only)] noncomputable instance category_as_type : small_category (as_type α) := { hom := λ i j, fin (fintype.card (@quiver.hom (obj_as_type α) _ i j)), id := λ i, fintype.equiv_fin _ (𝟙 i), comp := λ i j k f g, fintype.equiv_fin _ ((fintype.equiv_fin _).symm f ≫ (fintype.equiv_fin _).symm g) } -local attribute [simp] category_as_type_hom category_as_type_id - category_as_type_comp +local attribute [simp] category_as_type_id category_as_type_comp /-- The "identity" functor from `as_type α` to `obj_as_type α`. -/ @[simps] noncomputable def as_type_to_obj_as_type : as_type α ⥤ obj_as_type α := diff --git a/src/category_theory/groupoid.lean b/src/category_theory/groupoid.lean index 8f6b2105beea3..4889ac68b3223 100644 --- a/src/category_theory/groupoid.lean +++ b/src/category_theory/groupoid.lean @@ -45,6 +45,11 @@ class groupoid (obj : Type u) extends category.{v} obj : Type (max u (v+1)) := restate_axiom groupoid.inv_comp' restate_axiom groupoid.comp_inv' + +initialize_simps_projections groupoid (-to_category_to_category_struct_to_quiver_hom, + to_category_to_category_struct_comp → comp, to_category_to_category_struct_id → id, + -to_category_to_category_struct, -to_category) + /-- A `large_groupoid` is a groupoid where the objects live in `Type (u+1)` while the morphisms live in `Type u`. diff --git a/src/category_theory/monoidal/braided.lean b/src/category_theory/monoidal/braided.lean index b46afa9e8ca1d..79d3d2935f8c2 100644 --- a/src/category_theory/monoidal/braided.lean +++ b/src/category_theory/monoidal/braided.lean @@ -221,6 +221,9 @@ class symmetric_category (C : Type u) [category.{v} C] [monoidal_category.{v} C] restate_axiom symmetric_category.symmetry' attribute [simp,reassoc] symmetric_category.symmetry +initialize_simps_projections symmetric_category + (to_braided_category_braiding → braiding, -to_braided_category) + variables (C : Type u₁) [category.{v₁} C] [monoidal_category C] [braided_category C] variables (D : Type u₂) [category.{v₂} D] [monoidal_category D] [braided_category D] variables (E : Type u₃) [category.{v₃} E] [monoidal_category E] [braided_category E] diff --git a/src/category_theory/sites/sheaf.lean b/src/category_theory/sites/sheaf.lean index 149a0b9049837..04a1ec67fd722 100644 --- a/src/category_theory/sites/sheaf.lean +++ b/src/category_theory/sites/sheaf.lean @@ -373,7 +373,7 @@ instance : has_add (P ⟶ Q) := instance : add_comm_group (P ⟶ Q) := function.injective.add_comm_group (λ (f : Sheaf.hom P Q), f.1) (λ _ _ h, Sheaf.hom.ext _ _ h) rfl (λ _ _, rfl) (λ _, rfl) (λ _ _, rfl) - (λ _ _, by { dsimp at *, ext, simpa [*] }) (λ _ _, by { dsimp at *, ext, simpa [*] }) + (λ _ _, by { ext, simpa [*] }) (λ _ _, by { ext, simpa [*] }) instance : preadditive (Sheaf J A) := { hom_group := λ P Q, infer_instance, From 8535b76e601f11868af3e612fbecb730998a5631 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sat, 8 Apr 2023 13:17:57 +0000 Subject: [PATCH 0794/1291] feat(linear_algebra/matrix/rank): generalize to rings (#18748) This addresses a TODO comment. To achieve this we generalize `submodule.finrank_le`, `submodule.finrank_quotient_le`, and `finrank_le_finrank_of_injective` from vector spaces to free modules, and add a new lemma `linear_map.finrank_range_le`. --- src/data/matrix/rank.lean | 68 +++++++++++-------- src/linear_algebra/finite_dimensional.lean | 18 ----- src/linear_algebra/finrank.lean | 6 ++ .../free_module/finite/rank.lean | 32 +++++++++ 4 files changed, 77 insertions(+), 47 deletions(-) diff --git a/src/data/matrix/rank.lean b/src/data/matrix/rank.lean index 2bc99be10f82d..f0aa5b5e58642 100644 --- a/src/data/matrix/rank.lean +++ b/src/data/matrix/rank.lean @@ -20,7 +20,6 @@ This definition does not depend on the choice of basis, see `matrix.rank_eq_finr ## TODO * Show that `matrix.rank` is equal to the row-rank and column-rank -* Generalize away from fields -/ @@ -30,44 +29,49 @@ namespace matrix open finite_dimensional -variables {m n o K : Type*} [m_fin : fintype m] [fintype n] [fintype o] -variables [decidable_eq n] [decidable_eq o] [field K] -variables (A : matrix m n K) +variables {m n o R : Type*} [m_fin : fintype m] [fintype n] [fintype o] +variables [decidable_eq n] [decidable_eq o] [comm_ring R] /-- The rank of a matrix is the rank of its image. -/ -noncomputable def rank : ℕ := finrank K A.to_lin'.range +noncomputable def rank (A : matrix m n R) : ℕ := finrank R A.to_lin'.range -@[simp] lemma rank_one : rank (1 : matrix n n K) = fintype.card n := +@[simp] lemma rank_one [strong_rank_condition R] : rank (1 : matrix n n R) = fintype.card n := by rw [rank, to_lin'_one, linear_map.range_id, finrank_top, finrank_pi] -@[simp] lemma rank_zero : rank (0 : matrix n n K) = 0 := +@[simp] lemma rank_zero [nontrivial R] : rank (0 : matrix m n R) = 0 := by rw [rank, linear_equiv.map_zero, linear_map.range_zero, finrank_bot] -lemma rank_le_card_width : A.rank ≤ fintype.card n := +lemma rank_le_card_width [strong_rank_condition R] (A : matrix m n R) : A.rank ≤ fintype.card n := begin - convert le_of_add_le_left (A.to_lin'.finrank_range_add_finrank_ker).le, - exact (finrank_pi K).symm, + haveI : module.finite R (n → R) := module.finite.pi, + haveI : module.free R (n → R) := module.free.pi _ _, + exact A.to_lin'.finrank_range_le.trans_eq (finrank_pi _) end -lemma rank_le_width {m n : ℕ} (A : matrix (fin m) (fin n) K) : A.rank ≤ n := +lemma rank_le_width [strong_rank_condition R] {m n : ℕ} (A : matrix (fin m) (fin n) R) : + A.rank ≤ n := A.rank_le_card_width.trans $ (fintype.card_fin n).le -lemma rank_mul_le (B : matrix n o K) : (A ⬝ B).rank ≤ A.rank := +lemma rank_mul_le [strong_rank_condition R] (A : matrix m n R) (B : matrix n o R) : + (A ⬝ B).rank ≤ A.rank := begin - refine linear_map.finrank_le_finrank_of_injective (submodule.of_le_injective _), - rw [to_lin'_mul], - exact linear_map.range_comp_le_range _ _, + rw [rank, rank, to_lin'_mul], + refine cardinal.to_nat_le_of_le_of_lt_aleph_0 _ (linear_map.rank_comp_le1 _ _), + rw [←cardinal.lift_lt_aleph_0], + have := lift_rank_range_le A.to_lin', + rw [rank_fun', cardinal.lift_nat_cast] at this, + exact this.trans_lt (cardinal.nat_lt_aleph_0 (fintype.card n)), end -lemma rank_unit (A : (matrix n n K)ˣ) : - (A : matrix n n K).rank = fintype.card n := +lemma rank_unit [strong_rank_condition R] (A : (matrix n n R)ˣ) : + (A : matrix n n R).rank = fintype.card n := begin refine le_antisymm (rank_le_card_width A) _, - have := rank_mul_le (A : matrix n n K) (↑A⁻¹ : matrix n n K), + have := rank_mul_le (A : matrix n n R) (↑A⁻¹ : matrix n n R), rwa [← mul_eq_mul, ← units.coe_mul, mul_inv_self, units.coe_one, rank_one] at this, end -lemma rank_of_is_unit (A : matrix n n K) (h : is_unit A) : +lemma rank_of_is_unit [strong_rank_condition R] (A : matrix n n R) (h : is_unit A) : A.rank = fintype.card n := by { obtain ⟨A, rfl⟩ := h, exact rank_unit A } @@ -75,19 +79,19 @@ include m_fin lemma rank_eq_finrank_range_to_lin {M₁ M₂ : Type*} [add_comm_group M₁] [add_comm_group M₂] - [module K M₁] [module K M₂] (v₁ : basis m K M₁) (v₂ : basis n K M₂) : - A.rank = finrank K (to_lin v₂ v₁ A).range := + [module R M₁] [module R M₂] (A : matrix m n R) (v₁ : basis m R M₁) (v₂ : basis n R M₂) : + A.rank = finrank R (to_lin v₂ v₁ A).range := begin - let e₁ := (pi.basis_fun K m).equiv v₁ (equiv.refl _), - let e₂ := (pi.basis_fun K n).equiv v₂ (equiv.refl _), - have range_e₂ : (e₂ : (n → K) →ₗ[K] M₂).range = ⊤, + let e₁ := (pi.basis_fun R m).equiv v₁ (equiv.refl _), + let e₂ := (pi.basis_fun R n).equiv v₂ (equiv.refl _), + have range_e₂ : (e₂ : (n → R) →ₗ[R] M₂).range = ⊤, { rw linear_map.range_eq_top, exact e₂.surjective }, refine linear_equiv.finrank_eq (e₁.of_submodules _ _ _), rw [← linear_map.range_comp, ← linear_map.range_comp_of_range_eq_top (to_lin v₂ v₁ A) range_e₂], congr' 1, apply linear_map.pi_ext', rintro i, apply linear_map.ext_ring, - have aux₁ := to_lin_self (pi.basis_fun K n) (pi.basis_fun K m) A i, - have aux₂ := basis.equiv_apply (pi.basis_fun K n) i v₂, + have aux₁ := to_lin_self (pi.basis_fun R n) (pi.basis_fun R m) A i, + have aux₂ := basis.equiv_apply (pi.basis_fun R n) i v₂, rw [to_lin_eq_to_lin'] at aux₁, rw [pi.basis_fun_apply, linear_map.coe_std_basis] at aux₁ aux₂, simp only [linear_map.comp_apply, e₁, e₂, linear_equiv.coe_coe, equiv.refl_apply, aux₁, aux₂, @@ -95,12 +99,18 @@ begin basis.equiv_apply], end -lemma rank_le_card_height : A.rank ≤ fintype.card m := -(submodule.finrank_le _).trans (finrank_pi K).le +lemma rank_le_card_height [strong_rank_condition R] (A : matrix m n R) : + A.rank ≤ fintype.card m := +begin + haveI : module.finite R (m → R) := module.finite.pi, + haveI : module.free R (m → R) := module.free.pi _ _, + exact (submodule.finrank_le _).trans (finrank_pi R).le +end omit m_fin -lemma rank_le_height {m n : ℕ} (A : matrix (fin m) (fin n) K) : A.rank ≤ m := +lemma rank_le_height [strong_rank_condition R] {m n : ℕ} (A : matrix (fin m) (fin n) R) : + A.rank ≤ m := A.rank_le_card_height.trans $ (fintype.card_fin m).le end matrix diff --git a/src/linear_algebra/finite_dimensional.lean b/src/linear_algebra/finite_dimensional.lean index 8caa119dccdbe..285285d308288 100644 --- a/src/linear_algebra/finite_dimensional.lean +++ b/src/linear_algebra/finite_dimensional.lean @@ -696,17 +696,6 @@ begin exact submodule.finite_dimensional_finset_sup _ _, 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_rank] using - s.subtype.rank_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_rank] using - (mkq s).rank_le_of_surjective (surjective_quot_mk _) - /-- 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) : @@ -1014,13 +1003,6 @@ lemma ker_eq_bot_iff_range_eq_top_of_finrank_eq_finrank [finite_dimensional K V] f.ker = ⊥ ↔ f.range = ⊤ := by rw [range_eq_top, ker_eq_bot, injective_iff_surjective_of_finrank_eq_finrank H] -theorem finrank_le_finrank_of_injective [finite_dimensional K V] [finite_dimensional K V₂] - {f : V →ₗ[K] V₂} (hf : function.injective f) : finrank K V ≤ finrank K V₂ := -calc finrank K V - = finrank K f.range + finrank K f.ker : (finrank_range_add_finrank_ker f).symm -... = finrank K f.range : by rw [ker_eq_bot.2 hf, finrank_bot, add_zero] -... ≤ finrank K V₂ : submodule.finrank_le _ - /-- Given a linear map `f` between two vector spaces with the same dimension, if `ker f = ⊥` then `linear_equiv_of_injective` is the induced isomorphism between the two vector spaces. -/ diff --git a/src/linear_algebra/finrank.lean b/src/linear_algebra/finrank.lean index baf0f4822cd01..b78c2a2f8dfdd 100644 --- a/src/linear_algebra/finrank.lean +++ b/src/linear_algebra/finrank.lean @@ -85,6 +85,12 @@ begin exact n.zero_le }, end +lemma finrank_le_finrank_of_rank_le_rank + (h : lift.{v'} (module.rank K V) ≤ cardinal.lift.{v} (module.rank K V₂)) + (h' : module.rank K V₂ < ℵ₀) : + finrank K V ≤ finrank K V₂ := +by simpa only [to_nat_lift] using to_nat_le_of_le_of_lt_aleph_0 (lift_lt_aleph_0.mpr h') h + section variables [nontrivial K] [no_zero_smul_divisors K V] diff --git a/src/linear_algebra/free_module/finite/rank.lean b/src/linear_algebra/free_module/finite/rank.lean index dfcfef1346576..c3567b67b0b6d 100644 --- a/src/linear_algebra/free_module/finite/rank.lean +++ b/src/linear_algebra/free_module/finite/rank.lean @@ -111,3 +111,35 @@ by { simp [finrank] } end comm_ring end finite_dimensional + +section +open finite_dimensional + +variables {R M N} +variables [ring R] [strong_rank_condition R] +variables [add_comm_group M] [module R M] +variables [add_comm_group N] [module R N] + +lemma linear_map.finrank_le_finrank_of_injective + [module.free R N] [module.finite R N] {f : M →ₗ[R] N} (hf : function.injective f) : + finrank R M ≤ finrank R N := +finrank_le_finrank_of_rank_le_rank + (linear_map.lift_rank_le_of_injective _ hf) (rank_lt_aleph_0 _ _) + +lemma linear_map.finrank_range_le [module.free R M] [module.finite R M] (f : M →ₗ[R] N) : + finrank R f.range ≤ finrank R M := +finrank_le_finrank_of_rank_le_rank (lift_rank_range_le f) (rank_lt_aleph_0 _ _) + +/-- The dimension of a submodule is bounded by the dimension of the ambient space. -/ +lemma submodule.finrank_le [module.free R M] [module.finite R M] (s : submodule R M) : + finrank R s ≤ finrank R M := +by simpa only [cardinal.to_nat_lift] using to_nat_le_of_le_of_lt_aleph_0 + (rank_lt_aleph_0 _ _) (rank_submodule_le s) + +/-- The dimension of a quotient is bounded by the dimension of the ambient space. -/ +lemma submodule.finrank_quotient_le [module.free R M] [module.finite R M] (s : submodule R M) : + finrank R (M ⧸ s) ≤ finrank R M := +by simpa only [cardinal.to_nat_lift] using to_nat_le_of_le_of_lt_aleph_0 + (rank_lt_aleph_0 _ _) ((submodule.mkq s).rank_le_of_surjective (surjective_quot_mk _)) + +end From 19cb3751e5e9b3d97adb51023949c50c13b5fdfd Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Sat, 8 Apr 2023 13:17:58 +0000 Subject: [PATCH 0795/1291] chore(*): add mathlib4 synchronization comments (#18753) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebraic_topology.simplex_category` * `analysis.hofer` * `analysis.normed.group.pointwise` * `analysis.specific_limits.basic` * `category_theory.idempotents.basic` * `category_theory.idempotents.functor_categories` * `category_theory.idempotents.functor_extension` * `category_theory.idempotents.karoubi` * `category_theory.idempotents.karoubi_karoubi` * `data.matrix.pequiv` * `data.nat.choose.multinomial` * `group_theory.perm.fin` * `linear_algebra.affine_space.combination` * `linear_algebra.finsupp_vector_space` * `linear_algebra.free_module.basic` * `linear_algebra.matrix.dot_product` * `linear_algebra.multilinear.basic` * `linear_algebra.multilinear.basis` * `linear_algebra.multilinear.tensor_product` * `linear_algebra.std_basis` * `linear_algebra.tensor_product_basis` * `ring_theory.mv_polynomial.basic` * `ring_theory.quotient_nilpotent` * `ring_theory.quotient_noetherian` * `ring_theory.simple_module` * `ring_theory.valuation.quotient` * `set_theory.zfc.basic` * `topology.algebra.filter_basis` * `topology.algebra.module.linear_pmap` * `topology.algebra.module.simple` * `topology.algebra.star_subalgebra` * `topology.algebra.uniform_field` * `topology.algebra.uniform_filter_basis` * `topology.metric_space.hausdorff_distance` --- src/algebraic_topology/simplex_category.lean | 3 +++ src/analysis/hofer.lean | 3 +++ src/analysis/normed/group/pointwise.lean | 3 +++ src/analysis/specific_limits/basic.lean | 3 +++ src/category_theory/idempotents/basic.lean | 3 +++ src/category_theory/idempotents/functor_categories.lean | 3 +++ src/category_theory/idempotents/functor_extension.lean | 3 +++ src/category_theory/idempotents/karoubi.lean | 3 +++ src/category_theory/idempotents/karoubi_karoubi.lean | 3 +++ src/data/matrix/pequiv.lean | 3 +++ src/data/nat/choose/multinomial.lean | 3 +++ src/group_theory/perm/fin.lean | 3 +++ src/linear_algebra/affine_space/combination.lean | 3 +++ src/linear_algebra/finsupp_vector_space.lean | 3 +++ src/linear_algebra/free_module/basic.lean | 3 +++ src/linear_algebra/matrix/dot_product.lean | 3 +++ src/linear_algebra/multilinear/basic.lean | 3 +++ src/linear_algebra/multilinear/basis.lean | 3 +++ src/linear_algebra/multilinear/tensor_product.lean | 3 +++ src/linear_algebra/std_basis.lean | 3 +++ src/linear_algebra/tensor_product_basis.lean | 3 +++ src/ring_theory/mv_polynomial/basic.lean | 3 +++ src/ring_theory/quotient_nilpotent.lean | 3 +++ src/ring_theory/quotient_noetherian.lean | 3 +++ src/ring_theory/simple_module.lean | 3 +++ src/ring_theory/valuation/quotient.lean | 3 +++ src/set_theory/zfc/basic.lean | 3 +++ src/topology/algebra/filter_basis.lean | 3 +++ src/topology/algebra/module/linear_pmap.lean | 3 +++ src/topology/algebra/module/simple.lean | 3 +++ src/topology/algebra/star_subalgebra.lean | 3 +++ src/topology/algebra/uniform_field.lean | 3 +++ src/topology/algebra/uniform_filter_basis.lean | 3 +++ src/topology/metric_space/hausdorff_distance.lean | 3 +++ 34 files changed, 102 insertions(+) diff --git a/src/algebraic_topology/simplex_category.lean b/src/algebraic_topology/simplex_category.lean index f331834ebe631..a75189e9c9056 100644 --- a/src/algebraic_topology/simplex_category.lean +++ b/src/algebraic_topology/simplex_category.lean @@ -12,6 +12,9 @@ import category_theory.functor.reflects_isomorphisms /-! # The simplex category +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We construct a skeletal model of the simplex category, with objects `ℕ` and the morphism `n ⟶ m` being the monotone maps from `fin (n+1)` to `fin (m+1)`. diff --git a/src/analysis/hofer.lean b/src/analysis/hofer.lean index eee0d642da7ed..990eebfa14f74 100644 --- a/src/analysis/hofer.lean +++ b/src/analysis/hofer.lean @@ -8,6 +8,9 @@ import analysis.specific_limits.basic /-! # Hofer's lemma +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This is an elementary lemma about complete metric spaces. It is motivated by an application to the bubbling-off analysis for holomorphic curves in symplectic topology. We are *very* far away from having these applications, but the proof here is a nice diff --git a/src/analysis/normed/group/pointwise.lean b/src/analysis/normed/group/pointwise.lean index 75df633bfd403..672dd575ae904 100644 --- a/src/analysis/normed/group/pointwise.lean +++ b/src/analysis/normed/group/pointwise.lean @@ -9,6 +9,9 @@ import topology.metric_space.hausdorff_distance /-! # Properties of pointwise addition of sets in normed groups +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We explore the relationships between pointwise addition of sets in normed groups, and the norm. Notably, we show that the sum of bounded sets remain bounded. -/ diff --git a/src/analysis/specific_limits/basic.lean b/src/analysis/specific_limits/basic.lean index 2230dbead1b92..aecf5fd432bf4 100644 --- a/src/analysis/specific_limits/basic.lean +++ b/src/analysis/specific_limits/basic.lean @@ -12,6 +12,9 @@ import topology.algebra.algebra /-! # A collection of specific limit computations +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file, by design, is independent of `normed_space` in the import hierarchy. It contains important specific limit computations in metric spaces, in ordered rings/fields, and in specific instances of these such as `ℝ`, `ℝ≥0` and `ℝ≥0∞`. diff --git a/src/category_theory/idempotents/basic.lean b/src/category_theory/idempotents/basic.lean index a19ef233af3e8..1efc764aa4533 100644 --- a/src/category_theory/idempotents/basic.lean +++ b/src/category_theory/idempotents/basic.lean @@ -9,6 +9,9 @@ import category_theory.abelian.basic /-! # Idempotent complete categories +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we define the notion of idempotent complete categories (also known as Karoubian categories, or pseudoabelian in the case of preadditive categories). diff --git a/src/category_theory/idempotents/functor_categories.lean b/src/category_theory/idempotents/functor_categories.lean index ca71ecfa055a6..3df7fa5b83120 100644 --- a/src/category_theory/idempotents/functor_categories.lean +++ b/src/category_theory/idempotents/functor_categories.lean @@ -9,6 +9,9 @@ import category_theory.idempotents.karoubi /-! # Idempotent completeness and functor categories +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define an instance `functor_category_is_idempotent_complete` expressing that a functor category `J ⥤ C` is idempotent complete when the target category `C` is. diff --git a/src/category_theory/idempotents/functor_extension.lean b/src/category_theory/idempotents/functor_extension.lean index 2178e4be73b4c..0307fbd5ed354 100644 --- a/src/category_theory/idempotents/functor_extension.lean +++ b/src/category_theory/idempotents/functor_extension.lean @@ -9,6 +9,9 @@ import category_theory.idempotents.karoubi /-! # Extension of functors to the idempotent completion +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we construct an extension `functor_extension₁` of functors `C ⥤ karoubi D` to functors `karoubi C ⥤ karoubi D`. This results in an equivalence `karoubi_universal₁ C D : (C ⥤ karoubi D) ≌ (karoubi C ⥤ karoubi D)`. diff --git a/src/category_theory/idempotents/karoubi.lean b/src/category_theory/idempotents/karoubi.lean index 16fc6fd644f28..7913e1a3cc043 100644 --- a/src/category_theory/idempotents/karoubi.lean +++ b/src/category_theory/idempotents/karoubi.lean @@ -11,6 +11,9 @@ import category_theory.equivalence /-! # The Karoubi envelope of a category +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we define the Karoubi envelope `karoubi C` of a category `C`. ## Main constructions and definitions diff --git a/src/category_theory/idempotents/karoubi_karoubi.lean b/src/category_theory/idempotents/karoubi_karoubi.lean index 12d9e69c2588b..f2afef9d26a62 100644 --- a/src/category_theory/idempotents/karoubi_karoubi.lean +++ b/src/category_theory/idempotents/karoubi_karoubi.lean @@ -9,6 +9,9 @@ import category_theory.idempotents.karoubi /-! # Idempotence of the Karoubi envelope +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we construct the equivalence of categories `karoubi_karoubi.equivalence C : karoubi C ≌ karoubi (karoubi C)` for any category `C`. diff --git a/src/data/matrix/pequiv.lean b/src/data/matrix/pequiv.lean index e2fa164cd32df..5aa86acc0a02b 100644 --- a/src/data/matrix/pequiv.lean +++ b/src/data/matrix/pequiv.lean @@ -9,6 +9,9 @@ import data.pequiv /-! # partial equivalences for matrices +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Using partial equivalences to represent matrices. This file introduces the function `pequiv.to_matrix`, which returns a matrix containing ones and zeros. For any partial equivalence `f`, `f.to_matrix i j = 1 ↔ f i = some j`. diff --git a/src/data/nat/choose/multinomial.lean b/src/data/nat/choose/multinomial.lean index 1670c3605d621..e3fb99386cb6c 100644 --- a/src/data/nat/choose/multinomial.lean +++ b/src/data/nat/choose/multinomial.lean @@ -15,6 +15,9 @@ import data.finsupp.multiset /-! # Multinomial +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the multinomial coefficient and several small lemma's for manipulating it. ## Main declarations diff --git a/src/group_theory/perm/fin.lean b/src/group_theory/perm/fin.lean index a80671f99906b..57ce537a6db18 100644 --- a/src/group_theory/perm/fin.lean +++ b/src/group_theory/perm/fin.lean @@ -10,6 +10,9 @@ import logic.equiv.fintype /-! # Permutations of `fin n` + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ open equiv diff --git a/src/linear_algebra/affine_space/combination.lean b/src/linear_algebra/affine_space/combination.lean index 41c118950e3ff..119e009c1ea41 100644 --- a/src/linear_algebra/affine_space/combination.lean +++ b/src/linear_algebra/affine_space/combination.lean @@ -15,6 +15,9 @@ import tactic.fin_cases /-! # Affine combinations of points +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines affine combinations of points. ## Main definitions diff --git a/src/linear_algebra/finsupp_vector_space.lean b/src/linear_algebra/finsupp_vector_space.lean index 82d263f0d7a13..30f725e375198 100644 --- a/src/linear_algebra/finsupp_vector_space.lean +++ b/src/linear_algebra/finsupp_vector_space.lean @@ -9,6 +9,9 @@ import linear_algebra.std_basis /-! # Linear structures on function with finite support `ι →₀ M` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains results on the `R`-module structure on functions of finite support from a type `ι` to an `R`-module `M`, in particular in the case that `R` is a field. diff --git a/src/linear_algebra/free_module/basic.lean b/src/linear_algebra/free_module/basic.lean index ce8be77ced04b..da064229d8fdf 100644 --- a/src/linear_algebra/free_module/basic.lean +++ b/src/linear_algebra/free_module/basic.lean @@ -14,6 +14,9 @@ import linear_algebra.tensor_product_basis # Free modules +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We introduce a class `module.free R M`, for `R` a `semiring` and `M` an `R`-module and we provide several basic instances for this class. diff --git a/src/linear_algebra/matrix/dot_product.lean b/src/linear_algebra/matrix/dot_product.lean index c6b916b3ff7eb..28b475fdb742c 100644 --- a/src/linear_algebra/matrix/dot_product.lean +++ b/src/linear_algebra/matrix/dot_product.lean @@ -10,6 +10,9 @@ import linear_algebra.std_basis /-! # Dot product of two vectors +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains some results on the map `matrix.dot_product`, which maps two vectors `v w : n → R` to the sum of the entrywise products `v i * w i`. diff --git a/src/linear_algebra/multilinear/basic.lean b/src/linear_algebra/multilinear/basic.lean index ce8c160726d99..bef802e1d31f6 100644 --- a/src/linear_algebra/multilinear/basic.lean +++ b/src/linear_algebra/multilinear/basic.lean @@ -14,6 +14,9 @@ import data.fintype.sort /-! # Multilinear maps +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define multilinear maps as maps from `Π(i : ι), M₁ i` to `M₂` which are linear in each coordinate. Here, `M₁ i` and `M₂` are modules over a ring `R`, and `ι` is an arbitrary type (although some statements will require it to be a fintype). This space, denoted by diff --git a/src/linear_algebra/multilinear/basis.lean b/src/linear_algebra/multilinear/basis.lean index 86a1f14dd9d4b..14f7477bac1da 100644 --- a/src/linear_algebra/multilinear/basis.lean +++ b/src/linear_algebra/multilinear/basis.lean @@ -9,6 +9,9 @@ import linear_algebra.multilinear.basic /-! # Multilinear maps in relation to bases. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves lemmas about the action of multilinear maps on basis vectors. ## TODO diff --git a/src/linear_algebra/multilinear/tensor_product.lean b/src/linear_algebra/multilinear/tensor_product.lean index 07a7bd39f336e..9047597e1e612 100644 --- a/src/linear_algebra/multilinear/tensor_product.lean +++ b/src/linear_algebra/multilinear/tensor_product.lean @@ -8,6 +8,9 @@ import linear_algebra.tensor_product /-! # Constructions relating multilinear maps and tensor products. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ namespace multilinear_map diff --git a/src/linear_algebra/std_basis.lean b/src/linear_algebra/std_basis.lean index c198c0bcd33be..0cc92187283bb 100644 --- a/src/linear_algebra/std_basis.lean +++ b/src/linear_algebra/std_basis.lean @@ -10,6 +10,9 @@ import linear_algebra.pi /-! # The standard basis +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the standard basis `pi.basis (s : ∀ j, basis (ι j) R (M j))`, which is the `Σ j, ι j`-indexed basis of Π j, M j`. The basis vectors are given by `pi.basis s ⟨j, i⟩ j' = linear_map.std_basis R M j' (s j) i = if j = j' then s i else 0`. diff --git a/src/linear_algebra/tensor_product_basis.lean b/src/linear_algebra/tensor_product_basis.lean index b25b27a330581..cc422c6c209ab 100644 --- a/src/linear_algebra/tensor_product_basis.lean +++ b/src/linear_algebra/tensor_product_basis.lean @@ -9,6 +9,9 @@ import linear_algebra.finsupp_vector_space /-! # Bases and dimensionality of tensor products of modules +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + These can not go into `linear_algebra.tensor_product` since they depend on `linear_algebra.finsupp_vector_space` which in turn imports `linear_algebra.tensor_product`. diff --git a/src/ring_theory/mv_polynomial/basic.lean b/src/ring_theory/mv_polynomial/basic.lean index 91b0e27c74232..597e65023f10f 100644 --- a/src/ring_theory/mv_polynomial/basic.lean +++ b/src/ring_theory/mv_polynomial/basic.lean @@ -12,6 +12,9 @@ import linear_algebra.finsupp_vector_space /-! # Multivariate polynomials over commutative rings +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains basic facts about multivariate polynomials over commutative rings, for example that the monomials form a basis. diff --git a/src/ring_theory/quotient_nilpotent.lean b/src/ring_theory/quotient_nilpotent.lean index 95eb145876376..221c44c4f72fd 100644 --- a/src/ring_theory/quotient_nilpotent.lean +++ b/src/ring_theory/quotient_nilpotent.lean @@ -8,6 +8,9 @@ import ring_theory.ideal.quotient_operations /-! # Nilpotent elements in quotient rings + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ lemma ideal.is_radical_iff_quotient_reduced {R : Type*} [comm_ring R] (I : ideal R) : diff --git a/src/ring_theory/quotient_noetherian.lean b/src/ring_theory/quotient_noetherian.lean index b430736ab1ee8..2bbef7744a475 100644 --- a/src/ring_theory/quotient_noetherian.lean +++ b/src/ring_theory/quotient_noetherian.lean @@ -8,6 +8,9 @@ import ring_theory.quotient_nilpotent /-! # Noetherian quotient rings and quotient modules + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ instance ideal.quotient.is_noetherian_ring {R : Type*} [comm_ring R] [h : is_noetherian_ring R] diff --git a/src/ring_theory/simple_module.lean b/src/ring_theory/simple_module.lean index 0130a60f002c5..1fee19fed5a38 100644 --- a/src/ring_theory/simple_module.lean +++ b/src/ring_theory/simple_module.lean @@ -9,6 +9,9 @@ import order.jordan_holder /-! # Simple Modules +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main Definitions * `is_simple_module` indicates that a module has no proper submodules (the only submodules are `⊥` and `⊤`). diff --git a/src/ring_theory/valuation/quotient.lean b/src/ring_theory/valuation/quotient.lean index dde2f37caf667..5b9839dab4fdf 100644 --- a/src/ring_theory/valuation/quotient.lean +++ b/src/ring_theory/valuation/quotient.lean @@ -10,6 +10,9 @@ import ring_theory.ideal.quotient_operations /-! # The valuation on a quotient ring +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The support of a valuation `v : valuation R Γ₀` is `supp v`. If `J` is an ideal of `R` with `h : J ⊆ supp v` then the induced valuation on R / J = `ideal.quotient J` is `on_quot v h`. diff --git a/src/set_theory/zfc/basic.lean b/src/set_theory/zfc/basic.lean index cb6a6be94c5e7..062030c1039ae 100644 --- a/src/set_theory/zfc/basic.lean +++ b/src/set_theory/zfc/basic.lean @@ -10,6 +10,9 @@ import order.well_founded /-! # A model of ZFC +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we model Zermelo-Fraenkel set theory (+ Choice) using Lean's underlying type theory. We do this in four main steps: * Define pre-sets inductively. diff --git a/src/topology/algebra/filter_basis.lean b/src/topology/algebra/filter_basis.lean index fe8f97746d553..8f9664a67dbda 100644 --- a/src/topology/algebra/filter_basis.lean +++ b/src/topology/algebra/filter_basis.lean @@ -9,6 +9,9 @@ import topology.algebra.module.basic /-! # Group and ring filter bases +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A `group_filter_basis` is a `filter_basis` on a group with some properties relating the basis to the group structure. The main theorem is that a `group_filter_basis` on a group gives a topology on the group which makes it into a topological group diff --git a/src/topology/algebra/module/linear_pmap.lean b/src/topology/algebra/module/linear_pmap.lean index 4580db88f328f..b36ab7fbcd01d 100644 --- a/src/topology/algebra/module/linear_pmap.lean +++ b/src/topology/algebra/module/linear_pmap.lean @@ -10,6 +10,9 @@ import topology.algebra.module.basic /-! # Partially defined linear operators over topological vector spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define basic notions of partially defined linear operators, which we call unbounded operators for short. In this file we prove all elementary properties of unbounded operators that do not assume that the diff --git a/src/topology/algebra/module/simple.lean b/src/topology/algebra/module/simple.lean index befba1c2b38dc..ad8dcaef9771e 100644 --- a/src/topology/algebra/module/simple.lean +++ b/src/topology/algebra/module/simple.lean @@ -9,6 +9,9 @@ import topology.algebra.module.basic /-! # The kernel of a linear function is closed or dense +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove (`linear_map.is_closed_or_dense_ker`) that the kernel of a linear function `f : M →ₗ[R] N` is either closed or dense in `M` provided that `N` is a simple module over `R`. This applies, e.g., to the case when `R = N` is a division ring. diff --git a/src/topology/algebra/star_subalgebra.lean b/src/topology/algebra/star_subalgebra.lean index de1d13425fc14..241c1da3132bd 100644 --- a/src/topology/algebra/star_subalgebra.lean +++ b/src/topology/algebra/star_subalgebra.lean @@ -10,6 +10,9 @@ import topology.algebra.star /-! # Topological star (sub)algebras +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A topological star algebra over a topological semiring `R` is a topological semiring with a compatible continuous scalar multiplication by elements of `R` and a continuous star operation. We reuse typeclass `has_continuous_smul` for topological algebras. diff --git a/src/topology/algebra/uniform_field.lean b/src/topology/algebra/uniform_field.lean index c1ac29242fc85..c074b36ce0502 100644 --- a/src/topology/algebra/uniform_field.lean +++ b/src/topology/algebra/uniform_field.lean @@ -10,6 +10,9 @@ import field_theory.subfield /-! # Completion of topological fields +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The goal of this file is to prove the main part of Proposition 7 of Bourbaki GT III 6.8 : The completion `hat K` of a Hausdorff topological field is a field if the image under diff --git a/src/topology/algebra/uniform_filter_basis.lean b/src/topology/algebra/uniform_filter_basis.lean index 1bafafce81721..58df1cd45c0b8 100644 --- a/src/topology/algebra/uniform_filter_basis.lean +++ b/src/topology/algebra/uniform_filter_basis.lean @@ -10,6 +10,9 @@ import topology.algebra.uniform_group /-! # Uniform properties of neighborhood bases in topological algebra +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This files contains properties of filter bases on algebraic structures that also require the theory of uniform spaces. diff --git a/src/topology/metric_space/hausdorff_distance.lean b/src/topology/metric_space/hausdorff_distance.lean index 3c498da5f248b..a0b022d137142 100644 --- a/src/topology/metric_space/hausdorff_distance.lean +++ b/src/topology/metric_space/hausdorff_distance.lean @@ -10,6 +10,9 @@ import topology.instances.ennreal /-! # Hausdorff distance +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The Hausdorff distance on subsets of a metric (or emetric) space. Given two subsets `s` and `t` of a metric space, their Hausdorff distance is the smallest `d` From 36172d6661e181c215108035483dbbeabd62942e Mon Sep 17 00:00:00 2001 From: themathqueen <2497250a@research.gla.ac.uk> Date: Sat, 8 Apr 2023 15:29:56 +0000 Subject: [PATCH 0796/1291] feat(analysis/inner_product_space/symmetric): Polarization for `is_R_or_C` (#18397) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR contains * polarization identity for `is_R_or_C` for `T.is_symmetric` * `T.is_symmetric` implies `(∀ x, ⟪x, T x⟫ = 0) ↔ T = 0` Co-authored-by: themathqueen <23701951+themathqueen@users.noreply.github.com> --- .../inner_product_space/symmetric.lean | 34 +++++++++++++++++++ 1 file changed, 34 insertions(+) diff --git a/src/analysis/inner_product_space/symmetric.lean b/src/analysis/inner_product_space/symmetric.lean index bd358edccdbbb..c86dd6988052b 100644 --- a/src/analysis/inner_product_space/symmetric.lean +++ b/src/analysis/inner_product_space/symmetric.lean @@ -154,4 +154,38 @@ end end complex +/-- Polarization identity for symmetric linear maps. +See `inner_map_polarization` for the complex version without the symmetric assumption. -/ +lemma is_symmetric.inner_map_polarization {T : E →ₗ[𝕜] E} (hT : T.is_symmetric) (x y : E) : + ⟪T x, y⟫ = (⟪T (x + y), x + y⟫ - ⟪T (x - y), x - y⟫ - + I * ⟪T (x + (I : 𝕜) • y), x + (I : 𝕜) • y⟫ + + I * ⟪T (x - (I : 𝕜) • y), x - (I : 𝕜) • y⟫) / 4 := +begin + rcases @I_mul_I_ax 𝕜 _ with (h | h), + { simp_rw [h, zero_mul, sub_zero, add_zero, map_add, map_sub, inner_add_left, + inner_add_right, inner_sub_left, inner_sub_right, hT x, ← inner_conj_symm x (T y)], + suffices : (re ⟪T y, x⟫ : 𝕜) = ⟪T y, x⟫, + { rw eq_conj_iff_re.mpr this, + ring_nf, }, + { rw ← re_add_im ⟪T y, x⟫, + simp_rw [h, mul_zero, add_zero], + norm_cast, } }, + { simp_rw [map_add, map_sub, inner_add_left, inner_add_right, inner_sub_left, inner_sub_right, + linear_map.map_smul, inner_smul_left, inner_smul_right, is_R_or_C.conj_I, mul_add, + mul_sub, sub_sub, ← mul_assoc, mul_neg, h, neg_neg, one_mul, neg_one_mul], + ring, }, +end + +/-- A symmetric linear map `T` is zero if and only if `⟪T x, x⟫_ℝ = 0` for all `x`. +See `inner_map_self_eq_zero` for the complex version without the symmetric assumption. -/ +lemma is_symmetric.inner_map_eq_zero {T : E →ₗ[𝕜] E} (hT : T.is_symmetric) : + (∀ x, ⟪T x, x⟫ = 0) ↔ T = 0 := +begin + simp_rw [linear_map.ext_iff, zero_apply], + refine ⟨λ h x, _, λ h, by simp_rw [h, inner_zero_left, forall_const]⟩, + rw [← @inner_self_eq_zero 𝕜, hT.inner_map_polarization], + simp_rw [h _], + ring, +end + end linear_map From 284fdd2962e67d2932fa3a79ce19fcf92d38e228 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sat, 8 Apr 2023 18:50:04 +0000 Subject: [PATCH 0797/1291] chore(linear_algebra/alternating,topology/algebra/module/multilinear): add a fun_like instance (#18766) Co-authored-by: Parcly Taxel --- src/analysis/analytic/inverse.lean | 2 +- src/analysis/calculus/cont_diff_def.lean | 3 +- src/analysis/normed_space/multilinear.lean | 6 ++-- src/linear_algebra/alternating.lean | 11 +++++-- src/topology/algebra/module/multilinear.lean | 32 ++++++++++++-------- 5 files changed, 33 insertions(+), 21 deletions(-) diff --git a/src/analysis/analytic/inverse.lean b/src/analysis/analytic/inverse.lean index 82fb6aaad1073..60268c17f70fd 100644 --- a/src/analysis/analytic/inverse.lean +++ b/src/analysis/analytic/inverse.lean @@ -170,7 +170,7 @@ begin { simp only [right_inv_coeff_one] }, simp only [right_inv, neg_inj], rw remove_zero_comp_of_pos _ _ (add_pos_of_nonneg_of_pos (n.zero_le) zero_lt_two), - congrm i.symm.to_continuous_linear_map.comp_continuous_multilinear_map (p.comp (λ k, _) _), + congr' 2 with k, by_cases hk : k < n+2; simp [hk, IH] end diff --git a/src/analysis/calculus/cont_diff_def.lean b/src/analysis/calculus/cont_diff_def.lean index dc54c755ea169..2dc7148f9e37b 100644 --- a/src/analysis/calculus/cont_diff_def.lean +++ b/src/analysis/calculus/cont_diff_def.lean @@ -571,7 +571,8 @@ begin change ((p' y 0) (init (@cons 0 (λ i, E) z 0))) (@cons 0 (λ i, E) z 0 (last 0)) = ((p' y 0) 0) z, unfold_coes, - congr }, + congr, + dec_trivial }, { convert (Hp'.mono (inter_subset_left v u)).congr (λ x hx, Hp'.zero_eq x hx.1), { ext x y, change p' x 0 (init (@snoc 0 (λ i : fin 1, E) 0 y)) y = p' x 0 0 y, diff --git a/src/analysis/normed_space/multilinear.lean b/src/analysis/normed_space/multilinear.lean index 86a981e9cc103..73546f022c1f3 100644 --- a/src/analysis/normed_space/multilinear.lean +++ b/src/analysis/normed_space/multilinear.lean @@ -820,7 +820,7 @@ variables {𝕜 ι} lemma mk_pi_field_apply_one_eq_self (f : continuous_multilinear_map 𝕜 (λ(i : ι), 𝕜) G) : continuous_multilinear_map.mk_pi_field 𝕜 ι (f (λi, 1)) = f := -to_multilinear_map_inj f.to_multilinear_map.mk_pi_ring_apply_one_eq_self +to_multilinear_map_injective f.to_multilinear_map.mk_pi_ring_apply_one_eq_self @[simp] lemma norm_mk_pi_field (z : G) : ‖continuous_multilinear_map.mk_pi_field 𝕜 ι z‖ = ‖z‖ := (multilinear_map.mk_continuous_norm_le _ (norm_nonneg z) _).antisymm $ @@ -830,7 +830,7 @@ lemma mk_pi_field_eq_iff {z₁ z₂ : G} : continuous_multilinear_map.mk_pi_field 𝕜 ι z₁ = continuous_multilinear_map.mk_pi_field 𝕜 ι z₂ ↔ z₁ = z₂ := begin - rw [← to_multilinear_map_inj.eq_iff], + rw [← to_multilinear_map_injective.eq_iff], exact multilinear_map.mk_pi_ring_eq_iff end @@ -1225,7 +1225,7 @@ end @[simp] lemma continuous_multilinear_map.uncurry_curry_left (f : continuous_multilinear_map 𝕜 Ei G) : f.curry_left.uncurry_left = f := -continuous_multilinear_map.to_multilinear_map_inj $ f.to_multilinear_map.uncurry_curry_left +continuous_multilinear_map.to_multilinear_map_injective $ f.to_multilinear_map.uncurry_curry_left variables (𝕜 Ei G) diff --git a/src/linear_algebra/alternating.lean b/src/linear_algebra/alternating.lean index 73f03c6f237e9..fdaee20a618eb 100644 --- a/src/linear_algebra/alternating.lean +++ b/src/linear_algebra/alternating.lean @@ -80,7 +80,12 @@ open function /-! Basic coercion simp lemmas, largely copied from `ring_hom` and `multilinear_map` -/ section coercions -instance : has_coe_to_fun (alternating_map R M N ι) (λ _, (ι → M) → N) := ⟨λ x, x.to_fun⟩ +instance fun_like : fun_like (alternating_map R M N ι) (ι → M) (λ _, N) := +{ coe := alternating_map.to_fun, + coe_injective' := λ f g h, by { cases f, cases g, congr' } } + +-- shortcut instance +instance : has_coe_to_fun (alternating_map R M N ι) (λ _, (ι → M) → N) := ⟨fun_like.coe⟩ initialize_simps_projections alternating_map (to_fun → apply) @@ -96,14 +101,14 @@ theorem congr_arg (f : alternating_map R M N ι) {x y : ι → M} (h : x = y) : congr_arg (λ x : ι → M, f x) h theorem coe_injective : injective (coe_fn : alternating_map R M N ι → ((ι → M) → N)) := -λ f g h, by { cases f, cases g, cases h, refl } +fun_like.coe_injective @[simp, norm_cast] theorem coe_inj {f g : alternating_map R M N ι} : (f : (ι → M) → N) = g ↔ f = g := coe_injective.eq_iff @[ext] theorem ext {f f' : alternating_map R M N ι} (H : ∀ x, f x = f' x) : f = f' := -coe_injective (funext H) +fun_like.ext _ _ H theorem ext_iff {f g : alternating_map R M N ι} : f = g ↔ ∀ x, f x = g x := ⟨λ h x, h ▸ rfl, λ h, ext h⟩ diff --git a/src/topology/algebra/module/multilinear.lean b/src/topology/algebra/module/multilinear.lean index 0559849b68edd..5aebcb46546db 100644 --- a/src/topology/algebra/module/multilinear.lean +++ b/src/topology/algebra/module/multilinear.lean @@ -63,8 +63,19 @@ variables [semiring R] [topological_space M₂] [topological_space M₃] [topological_space M₄] (f f' : continuous_multilinear_map R M₁ M₂) +theorem to_multilinear_map_injective : + function.injective (continuous_multilinear_map.to_multilinear_map : + continuous_multilinear_map R M₁ M₂ → multilinear_map R M₁ M₂) +| ⟨f, hf⟩ ⟨g, hg⟩ rfl := rfl + +instance continuous_map_class : + continuous_map_class (continuous_multilinear_map R M₁ M₂) (Π i, M₁ i) M₂ := +{ coe := λ f, f.to_fun, + coe_injective' := λ f g h, to_multilinear_map_injective $ multilinear_map.coe_injective h, + map_continuous := continuous_multilinear_map.cont } + instance : has_coe_to_fun (continuous_multilinear_map R M₁ M₂) (λ _, (Π i, M₁ i) → M₂) := -⟨λ f, f.to_fun⟩ +⟨λ f, f⟩ /-- See Note [custom simps projection]. We need to specify this projection explicitly in this case, because it is a composition of multiple projections. -/ @@ -77,16 +88,11 @@ initialize_simps_projections continuous_multilinear_map @[simp] lemma coe_coe : (f.to_multilinear_map : (Π i, M₁ i) → M₂) = f := rfl -theorem to_multilinear_map_inj : - function.injective (continuous_multilinear_map.to_multilinear_map : - continuous_multilinear_map R M₁ M₂ → multilinear_map R M₁ M₂) -| ⟨f, hf⟩ ⟨g, hg⟩ rfl := rfl - @[ext] theorem ext {f f' : continuous_multilinear_map R M₁ M₂} (H : ∀ x, f x = f' x) : f = f' := -to_multilinear_map_inj $ multilinear_map.ext H +fun_like.ext _ _ H theorem ext_iff {f f' : continuous_multilinear_map R M₁ M₂} : f = f' ↔ ∀ x, f x = f' x := -by rw [← to_multilinear_map_inj.eq_iff, multilinear_map.ext_iff]; refl +by rw [← to_multilinear_map_injective.eq_iff, multilinear_map.ext_iff]; refl @[simp] lemma map_add [decidable_eq ι] (m : Πi, M₁ i) (i : ι) (x y : M₁ i) : f (update m i (x + y)) = f (update m i x) + f (update m i y) := @@ -142,7 +148,7 @@ instance [distrib_mul_action R'ᵐᵒᵖ M₂] [is_central_scalar R' M₂] : ⟨λ c₁ f, ext $ λ x, op_smul_eq_smul _ _⟩ instance : mul_action R' (continuous_multilinear_map A M₁ M₂) := -function.injective.mul_action to_multilinear_map to_multilinear_map_inj (λ _ _, rfl) +function.injective.mul_action to_multilinear_map to_multilinear_map_injective (λ _ _, rfl) end has_smul @@ -159,7 +165,7 @@ instance : has_add (continuous_multilinear_map R M₁ M₂) := rfl instance add_comm_monoid : add_comm_monoid (continuous_multilinear_map R M₁ M₂) := -to_multilinear_map_inj.add_comm_monoid _ rfl (λ _ _, rfl) (λ _ _, rfl) +to_multilinear_map_injective.add_comm_monoid _ rfl (λ _ _, rfl) (λ _ _, rfl) /-- Evaluation of a `continuous_multilinear_map` at a vector as an `add_monoid_hom`. -/ def apply_add_hom (m : Π i, M₁ i) : continuous_multilinear_map R M₁ M₂ →+ M₂ := @@ -359,7 +365,7 @@ instance : has_sub (continuous_multilinear_map R M₁ M₂) := @[simp] lemma sub_apply (m : Πi, M₁ i) : (f - f') m = f m - f' m := rfl instance : add_comm_group (continuous_multilinear_map R M₁ M₂) := -to_multilinear_map_inj.add_comm_group _ +to_multilinear_map_injective.add_comm_group _ rfl (λ _ _, rfl) (λ _, rfl) (λ _ _, rfl) (λ _ _, rfl) (λ _ _, rfl) end topological_add_group @@ -398,7 +404,7 @@ variables {R' R'' A : Type*} [monoid R'] [monoid R''] [semiring A] instance [has_continuous_add M₂] : distrib_mul_action R' (continuous_multilinear_map A M₁ M₂) := function.injective.distrib_mul_action ⟨to_multilinear_map, to_multilinear_map_zero, to_multilinear_map_add⟩ - to_multilinear_map_inj (λ _ _, rfl) + to_multilinear_map_injective (λ _ _, rfl) end distrib_mul_action @@ -414,7 +420,7 @@ variables {R' A : Type*} [semiring R'] [semiring A] pointwise addition and scalar multiplication. -/ instance : module R' (continuous_multilinear_map A M₁ M₂) := function.injective.module _ ⟨to_multilinear_map, to_multilinear_map_zero, to_multilinear_map_add⟩ - to_multilinear_map_inj (λ _ _, rfl) + to_multilinear_map_injective (λ _ _, rfl) /-- Linear map version of the map `to_multilinear_map` associating to a continuous multilinear map the corresponding multilinear map. -/ From 45ce3929e3bf9a086a216feea3b1ab6c14bf0e67 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sat, 8 Apr 2023 20:46:10 +0000 Subject: [PATCH 0798/1291] chore(linear_algebra/dimension): tidy lemma names for `linear_map.rank` (#18769) `le1` and `le2` don't really fit the usual naming convention. Also generalizes `rank_le_domain` to rings and `rank_le_range` across universes. --- src/data/matrix/rank.lean | 2 +- src/linear_algebra/dimension.lean | 12 ++++++------ src/linear_algebra/matrix/to_lin.lean | 5 +++-- 3 files changed, 10 insertions(+), 9 deletions(-) diff --git a/src/data/matrix/rank.lean b/src/data/matrix/rank.lean index f0aa5b5e58642..f112526ca70a7 100644 --- a/src/data/matrix/rank.lean +++ b/src/data/matrix/rank.lean @@ -56,7 +56,7 @@ lemma rank_mul_le [strong_rank_condition R] (A : matrix m n R) (B : matrix n o R (A ⬝ B).rank ≤ A.rank := begin rw [rank, rank, to_lin'_mul], - refine cardinal.to_nat_le_of_le_of_lt_aleph_0 _ (linear_map.rank_comp_le1 _ _), + refine cardinal.to_nat_le_of_le_of_lt_aleph_0 _ (linear_map.rank_comp_le_left _ _), rw [←cardinal.lift_lt_aleph_0], have := lift_rank_range_le A.to_lin', rw [rank_fun', cardinal.lift_nat_cast] at this, diff --git a/src/linear_algebra/dimension.lean b/src/linear_algebra/dimension.lean index 55283a585b87e..ee150002b15e5 100644 --- a/src/linear_algebra/dimension.lean +++ b/src/linear_algebra/dimension.lean @@ -1317,15 +1317,18 @@ variables [add_comm_group V'] [module K V'] /-- `rank f` is the rank of a `linear_map` `f`, defined as the dimension of `f.range`. -/ def rank (f : V →ₗ[K] V') : cardinal := module.rank K f.range -lemma rank_le_range (f : V →ₗ[K] V₁) : rank f ≤ module.rank K V₁ := +lemma rank_le_range (f : V →ₗ[K] V') : rank f ≤ module.rank K V' := rank_submodule_le _ +lemma rank_le_domain (f : V →ₗ[K] V₁) : rank f ≤ module.rank K V := +rank_range_le _ + @[simp] lemma rank_zero [nontrivial K] : rank (0 : V →ₗ[K] V') = 0 := by rw [rank, linear_map.range_zero, rank_bot] variables [add_comm_group V''] [module K V''] -lemma rank_comp_le1 (g : V →ₗ[K] V') (f : V' →ₗ[K] V'') : rank (f.comp g) ≤ rank f := +lemma rank_comp_le_left (g : V →ₗ[K] V') (f : V' →ₗ[K] V'') : rank (f.comp g) ≤ rank f := begin refine rank_le_of_submodule _ _ _, rw [linear_map.range_comp], @@ -1334,7 +1337,7 @@ end variables [add_comm_group V'₁] [module K V'₁] -lemma rank_comp_le2 (g : V →ₗ[K] V') (f : V' →ₗ[K] V'₁) : rank (f.comp g) ≤ rank g := +lemma rank_comp_le_right (g : V →ₗ[K] V') (f : V' →ₗ[K] V'₁) : rank (f.comp g) ≤ rank g := by rw [rank, rank, linear_map.range_comp]; exact rank_map_le _ _ end ring @@ -1343,9 +1346,6 @@ section division_ring variables [division_ring K] [add_comm_group V] [module K V] [add_comm_group V₁] [module K V₁] variables [add_comm_group V'] [module K V'] -lemma rank_le_domain (f : V →ₗ[K] V₁) : rank f ≤ module.rank K V := -by { rw [← rank_range_add_rank_ker f], exact self_le_add_right _ _ } - lemma rank_add_le (f g : V →ₗ[K] V') : rank (f + g) ≤ rank f + rank g := calc rank (f + g) ≤ module.rank K (f.range ⊔ g.range : submodule K V') : begin diff --git a/src/linear_algebra/matrix/to_lin.lean b/src/linear_algebra/matrix/to_lin.lean index 380fa4e5b4e41..8374debb11dce 100644 --- a/src/linear_algebra/matrix/to_lin.lean +++ b/src/linear_algebra/matrix/to_lin.lean @@ -342,12 +342,13 @@ lemma linear_map.to_matrix_alg_equiv'_mul (f * g).to_matrix_alg_equiv' = f.to_matrix_alg_equiv' ⬝ g.to_matrix_alg_equiv' := linear_map.to_matrix_alg_equiv'_comp f g -lemma matrix.rank_vec_mul_vec {K m n : Type u} [field K] [fintype n] [decidable_eq n] +lemma matrix.rank_vec_mul_vec {K m n : Type u} + [comm_ring K] [strong_rank_condition K] [fintype n] [decidable_eq n] (w : m → K) (v : n → K) : rank (vec_mul_vec w v).to_lin' ≤ 1 := begin rw [vec_mul_vec_eq, matrix.to_lin'_mul], - refine le_trans (rank_comp_le1 _ _) _, + refine le_trans (rank_comp_le_left _ _) _, refine (rank_le_domain _).trans_eq _, rw [rank_fun', fintype.card_unit, nat.cast_one] end From b1c23399f01266afe392a0d8f71f599a0dad4f7b Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sun, 9 Apr 2023 07:51:18 +0000 Subject: [PATCH 0799/1291] refactor(linear_algebra/matrix/finite_dimensional): deduplicate (#18770) * `matrix.finrank_matrix` was a duplicate of `finite_dimensional.finrank_matrix`. * `linear_map.finrank_linear_map` was a duplicate of `finrank_linear_hom`, now merged to `finite_dimensional.finrank_linear_map` * `finite_dimensional.linear_map` was a duplicate of `linear_map.finite_dimensional` and can be golfed using `module.finite.linear_map` * `finite_dimensional.matrix` can be golfed using `module.finite.matrix` For now, I've left behind `finite_dimensional` instances, but proved them in terms of the `module.finite` versions. To enable this, some imports have been adjusted. The resulting import structure substantially cuts the dependencies consumed by `linear_algebra.matrix.to_lin`; it no longer needs `module.rank` to be available. Co-authored-by: Jeremy Tan Jie Rui --- src/algebra/category/fgModule/basic.lean | 3 +- .../normed_space/finite_dimension.lean | 1 - src/field_theory/tower.lean | 17 +----- src/linear_algebra/bilinear_form.lean | 1 + src/linear_algebra/determinant.lean | 1 + .../free_module/finite/matrix.lean | 40 +++++++------ .../free_module/finite/rank.lean | 2 +- src/linear_algebra/matrix/diagonal.lean | 1 + .../matrix/finite_dimensional.lean | 40 ++++++++----- src/linear_algebra/matrix/to_lin.lean | 57 ------------------- .../matrix/to_linear_equiv.lean | 1 + .../algebra/module/finite_dimension.lean | 1 + 12 files changed, 59 insertions(+), 106 deletions(-) diff --git a/src/algebra/category/fgModule/basic.lean b/src/algebra/category/fgModule/basic.lean index 6ea72337a7024..c439cce1df2ee 100644 --- a/src/algebra/category/fgModule/basic.lean +++ b/src/algebra/category/fgModule/basic.lean @@ -6,6 +6,7 @@ Authors: Jakob von Raumer import category_theory.monoidal.rigid.basic import category_theory.monoidal.subcategory import linear_algebra.coevaluation +import linear_algebra.free_module.finite.matrix import algebra.category.Module.monoidal /-! @@ -131,7 +132,7 @@ instance (V W : fgModule K) : module.finite K (V ⟶ W) := instance closed_predicate_module_finite : monoidal_category.closed_predicate (λ V : Module.{u} K, module.finite K V) := -{ prop_ihom' := λ X Y hX hY, by exactI @linear_map.finite_dimensional K _ X _ _ hX Y _ _ hY } +{ prop_ihom' := λ X Y hX hY, by exactI @module.finite.linear_map K X Y _ _ _ _ _ _ _ hX hY } instance : monoidal_closed (fgModule K) := by dsimp_result { dsimp [fgModule], apply_instance, } diff --git a/src/analysis/normed_space/finite_dimension.lean b/src/analysis/normed_space/finite_dimension.lean index 0252f919cb9f8..29b0b0340edf1 100644 --- a/src/analysis/normed_space/finite_dimension.lean +++ b/src/analysis/normed_space/finite_dimension.lean @@ -8,7 +8,6 @@ import analysis.normed_space.add_torsor import analysis.normed_space.affine_isometry import analysis.normed_space.operator_norm import analysis.normed_space.riesz_lemma -import linear_algebra.matrix.to_lin import topology.algebra.module.finite_dimension import topology.algebra.infinite_sum.module import topology.instances.matrix diff --git a/src/field_theory/tower.lean b/src/field_theory/tower.lean index 0a4321cf1f790..8c67f3c076912 100644 --- a/src/field_theory/tower.lean +++ b/src/field_theory/tower.lean @@ -114,25 +114,10 @@ theorem subalgebra.is_simple_order_of_finrank_prime (A) [ring A] [is_domain A] [ end } /- TODO: `intermediate_field` version -/ -instance linear_map (F : Type u) (V : Type v) (W : Type w) - [field F] [add_comm_group V] [module F V] [add_comm_group W] [module F W] - [finite_dimensional F V] [finite_dimensional F W] : - finite_dimensional F (V →ₗ[F] W) := -let b := basis.of_vector_space F V, c := basis.of_vector_space F W in -(matrix.to_lin b c).finite_dimensional - -lemma finrank_linear_map (F : Type u) (V : Type v) (W : Type w) - [field F] [add_comm_group V] [module F V] [add_comm_group W] [module F W] - [finite_dimensional F V] [finite_dimensional F W] : - finrank F (V →ₗ[F] W) = finrank F V * finrank F W := - let b := basis.of_vector_space F V, c := basis.of_vector_space F W in -by rw [linear_equiv.finrank_eq (linear_map.to_matrix b c), matrix.finrank_matrix, - finrank_eq_card_basis b, finrank_eq_card_basis c, mul_comm] - -- TODO: generalize by removing [finite_dimensional F K] -- V = ⊕F, -- (V →ₗ[F] K) = ((⊕F) →ₗ[F] K) = (⊕ (F →ₗ[F] K)) = ⊕K -instance linear_map' (F : Type u) (K : Type v) (V : Type w) +instance _root_.linear_map.finite_dimensional'' (F : Type u) (K : Type v) (V : Type w) [field F] [field K] [algebra F K] [finite_dimensional F K] [add_comm_group V] [module F V] [finite_dimensional F V] : finite_dimensional K (V →ₗ[F] K) := diff --git a/src/linear_algebra/bilinear_form.lean b/src/linear_algebra/bilinear_form.lean index 68e85e21118e8..2f3aeab57068b 100644 --- a/src/linear_algebra/bilinear_form.lean +++ b/src/linear_algebra/bilinear_form.lean @@ -5,6 +5,7 @@ Authors: Andreas Swerdlow, Kexing Ying -/ import linear_algebra.dual +import linear_algebra.free_module.finite.matrix import linear_algebra.matrix.to_lin /-! diff --git a/src/linear_algebra/determinant.lean b/src/linear_algebra/determinant.lean index 8f4a7c9dbf6f8..c70296ec4e78e 100644 --- a/src/linear_algebra/determinant.lean +++ b/src/linear_algebra/determinant.lean @@ -3,6 +3,7 @@ Copyright (c) 2019 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Patrick Massot, Casper Putz, Anne Baanen -/ +import linear_algebra.finite_dimensional import linear_algebra.general_linear_group import linear_algebra.matrix.reindex import tactic.field_simp diff --git a/src/linear_algebra/free_module/finite/matrix.lean b/src/linear_algebra/free_module/finite/matrix.lean index 4883618aa831b..4f134820b27f9 100644 --- a/src/linear_algebra/free_module/finite/matrix.lean +++ b/src/linear_algebra/free_module/finite/matrix.lean @@ -5,7 +5,7 @@ Authors: Riccardo Brasca -/ import linear_algebra.finrank -import linear_algebra.free_module.finite.basic +import linear_algebra.free_module.finite.rank import linear_algebra.matrix.to_lin /-! @@ -25,25 +25,26 @@ universes u v w variables (R : Type u) (M : Type v) (N : Type w) -namespace module.free +open module.free (choose_basis) +open finite_dimensional (finrank) section comm_ring variables [comm_ring R] [add_comm_group M] [module R M] [module.free R M] variables [add_comm_group N] [module R N] [module.free R N] -instance linear_map [module.finite R M] [module.finite R N] : module.free R (M →ₗ[R] N) := +instance module.free.linear_map [module.finite R M] [module.finite R N] : + module.free R (M →ₗ[R] N) := begin casesI subsingleton_or_nontrivial R, { apply module.free.of_subsingleton' }, classical, - exact of_equiv - (linear_map.to_matrix (module.free.choose_basis R M) (module.free.choose_basis R N)).symm, + exact module.free.of_equiv (linear_map.to_matrix (choose_basis R M) (choose_basis R N)).symm, end variables {R} -instance _root_.module.finite.linear_map [module.finite R M] [module.finite R N] : +instance module.finite.linear_map [module.finite R M] [module.finite R N] : module.finite R (M →ₗ[R] N) := begin casesI subsingleton_or_nontrivial R, @@ -60,10 +61,10 @@ section integer variables [add_comm_group M] [module.finite ℤ M] [module.free ℤ M] variables [add_comm_group N] [module.finite ℤ N] [module.free ℤ N] -instance _root_.module.finite.add_monoid_hom : module.finite ℤ (M →+ N) := +instance module.finite.add_monoid_hom : module.finite ℤ (M →+ N) := module.finite.equiv (add_monoid_hom_lequiv_int ℤ).symm -instance add_monoid_hom : module.free ℤ (M →+ N) := +instance module.free.add_monoid_hom : module.free ℤ (M →+ N) := begin letI : module.free ℤ (M →ₗ[ℤ] N) := module.free.linear_map _ _ _, exact module.free.of_equiv (add_monoid_hom_lequiv_int ℤ).symm @@ -73,25 +74,30 @@ end integer section comm_ring -open finite_dimensional - variables [comm_ring R] [strong_rank_condition R] variables [add_comm_group M] [module R M] [module.free R M] [module.finite R M] variables [add_comm_group N] [module R N] [module.free R N] [module.finite R N] /-- The finrank of `M →ₗ[R] N` is `(finrank R M) * (finrank R N)`. -/ ---TODO: this should follow from `linear_equiv.finrank_eq`, that is over a field. -lemma finrank_linear_hom : finrank R (M →ₗ[R] N) = (finrank R M) * (finrank R N) := +lemma finite_dimensional.finrank_linear_map : + finrank R (M →ₗ[R] N) = (finrank R M) * (finrank R N) := begin classical, letI := nontrivial_of_invariant_basis_number R, have h := (linear_map.to_matrix (choose_basis R M) (choose_basis R N)), - let b := (matrix.std_basis _ _ _).map h.symm, - rw [finrank, rank_eq_card_basis b, ← cardinal.mk_fintype, cardinal.mk_to_nat_eq_card, finrank, - finrank, rank_eq_card_choose_basis_index, rank_eq_card_choose_basis_index, - cardinal.mk_to_nat_eq_card, cardinal.mk_to_nat_eq_card, fintype.card_prod, mul_comm] + simp_rw [h.finrank_eq, finite_dimensional.finrank_matrix, + finite_dimensional.finrank_eq_card_choose_basis_index, mul_comm], end end comm_ring -end module.free +lemma matrix.rank_vec_mul_vec {K m n : Type u} + [comm_ring K] [strong_rank_condition K] [fintype n] [decidable_eq n] + (w : m → K) (v : n → K) : + (matrix.vec_mul_vec w v).to_lin'.rank ≤ 1 := +begin + rw [matrix.vec_mul_vec_eq, matrix.to_lin'_mul], + refine le_trans (linear_map.rank_comp_le_left _ _) _, + refine (linear_map.rank_le_domain _).trans_eq _, + rw [rank_fun', fintype.card_unit, nat.cast_one] +end diff --git a/src/linear_algebra/free_module/finite/rank.lean b/src/linear_algebra/free_module/finite/rank.lean index c3567b67b0b6d..b4160c3f448e2 100644 --- a/src/linear_algebra/free_module/finite/rank.lean +++ b/src/linear_algebra/free_module/finite/rank.lean @@ -90,7 +90,7 @@ end /-- If `m` and `n` are `fintype`, the finrank of `m × n` matrices is `(fintype.card m) * (fintype.card n)`. -/ -lemma finrank_matrix (m n : Type v) [fintype m] [fintype n] : +lemma finrank_matrix (m n : Type*) [fintype m] [fintype n] : finrank R (matrix m n R) = (card m) * (card n) := by { simp [finrank] } diff --git a/src/linear_algebra/matrix/diagonal.lean b/src/linear_algebra/matrix/diagonal.lean index b46dce0d98a54..24af38dbb344e 100644 --- a/src/linear_algebra/matrix/diagonal.lean +++ b/src/linear_algebra/matrix/diagonal.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Patrick Massot, Casper Putz, Anne Baanen -/ import linear_algebra.matrix.to_lin +import linear_algebra.free_module.rank /-! # Diagonal matrices diff --git a/src/linear_algebra/matrix/finite_dimensional.lean b/src/linear_algebra/matrix/finite_dimensional.lean index 55796c37dfd9f..5931077a135e0 100644 --- a/src/linear_algebra/matrix/finite_dimensional.lean +++ b/src/linear_algebra/matrix/finite_dimensional.lean @@ -5,17 +5,20 @@ Authors: Johannes Hölzl, Patrick Massot, Casper Putz, Anne Baanen -/ import data.matrix.basic import linear_algebra.finite_dimensional +import linear_algebra.free_module.finite.matrix +import linear_algebra.matrix.to_lin /-! # The finite-dimensional space of matrices -This file shows that `m` by `n` matrices form a finite-dimensional space, -and proves the `finrank` of that space is equal to `card m * card n`. +This file shows that `m` by `n` matrices form a finite-dimensional space. +Note that this is proven more generally elsewhere over modules as `module.finite.matrix`; this file +exists only to provide an entry in the instance list for `finite_dimensional`. ## Main definitions * `matrix.finite_dimensional`: matrices form a finite dimensional vector space over a field `K` - * `matrix.finrank_matrix`: the `finrank` of `matrix m n R` is `card m * card n` + * `linear_map.finite_dimensional` ## Tags @@ -32,17 +35,28 @@ section finite_dimensional variables {m n : Type*} {R : Type v} [field R] instance [finite m] [finite n] : finite_dimensional R (matrix m n R) := -linear_equiv.finite_dimensional (linear_equiv.curry R m n) - -/-- -The dimension of the space of finite dimensional matrices -is the product of the number of rows and columns. --/ -@[simp] lemma finrank_matrix [fintype m] [fintype n] : - finite_dimensional.finrank R (matrix m n R) = fintype.card m * fintype.card n := -by rw [@linear_equiv.finrank_eq R (matrix m n R) _ _ _ _ _ _ (linear_equiv.curry R m n).symm, - finite_dimensional.finrank_fintype_fun_eq_card, fintype.card_prod] +module.finite.matrix end finite_dimensional end matrix + +namespace linear_map + +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 : finite_dimensional K (V →ₗ[K] W) := +module.finite.linear_map _ _ + +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, as 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 linear_map diff --git a/src/linear_algebra/matrix/to_lin.lean b/src/linear_algebra/matrix/to_lin.lean index 8374debb11dce..65524e0897a21 100644 --- a/src/linear_algebra/matrix/to_lin.lean +++ b/src/linear_algebra/matrix/to_lin.lean @@ -5,7 +5,6 @@ Authors: Johannes Hölzl, Patrick Massot, Casper Putz, Anne Baanen -/ import data.matrix.block import data.matrix.notation -import linear_algebra.matrix.finite_dimensional import linear_algebra.std_basis import ring_theory.algebra_tower import algebra.module.algebra @@ -342,17 +341,6 @@ lemma linear_map.to_matrix_alg_equiv'_mul (f * g).to_matrix_alg_equiv' = f.to_matrix_alg_equiv' ⬝ g.to_matrix_alg_equiv' := linear_map.to_matrix_alg_equiv'_comp f g -lemma matrix.rank_vec_mul_vec {K m n : Type u} - [comm_ring K] [strong_rank_condition K] [fintype n] [decidable_eq n] - (w : m → K) (v : n → K) : - rank (vec_mul_vec w v).to_lin' ≤ 1 := -begin - rw [vec_mul_vec_eq, matrix.to_lin'_mul], - refine le_trans (rank_comp_le_left _ _) _, - refine (rank_le_domain _).trans_eq _, - rw [rank_fun', fintype.card_unit, nat.cast_one] -end - end to_matrix' section to_matrix @@ -723,51 +711,6 @@ end lmul_tower end algebra -namespace linear_map - -section finite_dimensional - -open_locale classical - -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 : 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, as 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. --/ -@[simp] lemma finrank_linear_map : - finite_dimensional.finrank K (V →ₗ[K] W) = - (finite_dimensional.finrank K V) * (finite_dimensional.finrank K W) := -begin - let hbV := basis.of_vector_space K V, - let hbW := basis.of_vector_space K W, - rw [linear_equiv.finrank_eq (linear_map.to_matrix hbV hbW), matrix.finrank_matrix, - finite_dimensional.finrank_eq_card_basis hbV, finite_dimensional.finrank_eq_card_basis hbW, - mul_comm], -end - -end finite_dimensional -end linear_map - section variables {R : Type v} [comm_ring R] {n : Type*} [decidable_eq n] diff --git a/src/linear_algebra/matrix/to_linear_equiv.lean b/src/linear_algebra/matrix/to_linear_equiv.lean index a8f4601fb4f24..f4a54f3bd0f7f 100644 --- a/src/linear_algebra/matrix/to_linear_equiv.lean +++ b/src/linear_algebra/matrix/to_linear_equiv.lean @@ -3,6 +3,7 @@ Copyright (c) 2019 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Patrick Massot, Casper Putz, Anne Baanen -/ +import linear_algebra.finite_dimensional import linear_algebra.matrix.nondegenerate import linear_algebra.matrix.nonsingular_inverse import linear_algebra.matrix.to_lin diff --git a/src/topology/algebra/module/finite_dimension.lean b/src/topology/algebra/module/finite_dimension.lean index 3a7b1d21c7da2..ad3e32a2cbd95 100644 --- a/src/topology/algebra/module/finite_dimension.lean +++ b/src/topology/algebra/module/finite_dimension.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel, Anatole Dedecker -/ import analysis.locally_convex.balanced_core_hull +import linear_algebra.free_module.finite.matrix import topology.algebra.module.simple import topology.algebra.module.determinant From e05ead7993520a432bec94ac504842d90707ad63 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Sun, 9 Apr 2023 12:17:07 +0000 Subject: [PATCH 0800/1291] chore(set_theory/cardinal/basic): less awkward placement of theorems (#18771) Some of the new results on limit cardinals were awkwardly breaking up blocks of code with related theorems, probably due to some botched merge. We reorder them to avoid this. Ported along other changes to this file in: https://github.com/leanprover-community/mathlib4/pull/3343 --- src/set_theory/cardinal/basic.lean | 66 +++++++++++++++--------------- 1 file changed, 33 insertions(+), 33 deletions(-) diff --git a/src/set_theory/cardinal/basic.lean b/src/set_theory/cardinal/basic.lean index f0028ca8d6974..f7eb3ef0a70a5 100644 --- a/src/set_theory/cardinal/basic.lean +++ b/src/set_theory/cardinal/basic.lean @@ -561,21 +561,11 @@ instance : succ_order cardinal := succ_order.of_succ_le_iff (λ c, Inf {c' | c < c'}) (λ a b, ⟨lt_of_lt_of_le $ Inf_mem $ exists_gt a, cInf_le'⟩) -/-- A cardinal is a limit if it is not zero or a successor cardinal. Note that `ℵ₀` is a limit - cardinal by this definition, but `0` isn't. - - Use `is_succ_limit` if you want to include the `c = 0` case. -/ -def is_limit (c : cardinal) : Prop := c ≠ 0 ∧ is_succ_limit c - -protected theorem is_limit.ne_zero {c} (h : is_limit c) : c ≠ 0 := h.1 - -protected theorem is_limit.is_succ_limit {c} (h : is_limit c) : is_succ_limit c := h.2 - -theorem is_limit.succ_lt {x c} (h : is_limit c) : x < c → succ x < c := h.is_succ_limit.succ_lt +theorem succ_def (c : cardinal) : succ c = Inf {c' | c < c'} := rfl -theorem is_succ_limit_zero : is_succ_limit (0 : cardinal) := is_succ_limit_bot +lemma succ_pos : ∀ c : cardinal, 0 < succ c := bot_lt_succ -theorem succ_def (c : cardinal) : succ c = Inf {c' | c < c'} := rfl +lemma succ_ne_zero (c : cardinal) : succ c ≠ 0 := (succ_pos _).ne' theorem add_one_le_succ (c : cardinal.{u}) : c + 1 ≤ succ c := begin @@ -589,9 +579,19 @@ begin ... ≤ #β : (f.option_elim b hb).cardinal_le end -lemma succ_pos : ∀ c : cardinal, 0 < succ c := bot_lt_succ +/-- A cardinal is a limit if it is not zero or a successor cardinal. Note that `ℵ₀` is a limit + cardinal by this definition, but `0` isn't. -lemma succ_ne_zero (c : cardinal) : succ c ≠ 0 := (succ_pos _).ne' + Use `is_succ_limit` if you want to include the `c = 0` case. -/ +def is_limit (c : cardinal) : Prop := c ≠ 0 ∧ is_succ_limit c + +protected theorem is_limit.ne_zero {c} (h : is_limit c) : c ≠ 0 := h.1 + +protected theorem is_limit.is_succ_limit {c} (h : is_limit c) : is_succ_limit c := h.2 + +theorem is_limit.succ_lt {x c} (h : is_limit c) : x < c → succ x < c := h.is_succ_limit.succ_lt + +theorem is_succ_limit_zero : is_succ_limit (0 : cardinal) := is_succ_limit_bot /-- The indexed sum of cardinals is the cardinality of the indexed disjoint union, i.e. sigma type. -/ @@ -1015,6 +1015,24 @@ theorem aleph_0_le {c : cardinal} : ℵ₀ ≤ c ↔ ∀ n : ℕ, ↑n ≤ c := exact (nat.lt_succ_self _).not_le (nat_cast_le.1 (h (n+1))) end⟩ +theorem is_succ_limit_aleph_0 : is_succ_limit ℵ₀ := +is_succ_limit_of_succ_lt $ λ a ha, begin + rcases lt_aleph_0.1 ha with ⟨n, rfl⟩, + rw ←nat_succ, + apply nat_lt_aleph_0 +end + +theorem is_limit_aleph_0 : is_limit ℵ₀ := ⟨aleph_0_ne_zero, is_succ_limit_aleph_0⟩ + +theorem is_limit.aleph_0_le {c : cardinal} (h : is_limit c) : ℵ₀ ≤ c := +begin + by_contra' h', + rcases lt_aleph_0.1 h' with ⟨_ | n, rfl⟩, + { exact h.ne_zero.irrefl }, + { rw nat_succ at h, + exact not_is_succ_limit_succ _ h.is_succ_limit } +end + @[simp] lemma range_nat_cast : range (coe : ℕ → cardinal) = Iio ℵ₀ := ext $ λ x, by simp only [mem_Iio, mem_range, eq_comm, lt_aleph_0] @@ -1035,24 +1053,6 @@ lt_aleph_0_iff_finite.trans finite_coe_iff alias lt_aleph_0_iff_set_finite ↔ _ _root_.set.finite.lt_aleph_0 -theorem is_succ_limit_aleph_0 : is_succ_limit ℵ₀ := -is_succ_limit_of_succ_lt $ λ a ha, begin - rcases lt_aleph_0.1 ha with ⟨n, rfl⟩, - rw ←nat_succ, - apply nat_lt_aleph_0 -end - -theorem is_limit_aleph_0 : is_limit ℵ₀ := ⟨aleph_0_ne_zero, is_succ_limit_aleph_0⟩ - -theorem is_limit.aleph_0_le {c : cardinal} (h : is_limit c) : ℵ₀ ≤ c := -begin - by_contra' h', - rcases lt_aleph_0.1 h' with ⟨_ | n, rfl⟩, - { exact h.ne_zero.irrefl }, - { rw nat_succ at h, - exact not_is_succ_limit_succ _ h.is_succ_limit } -end - @[simp] theorem lt_aleph_0_iff_subtype_finite {p : α → Prop} : #{x // p x} < ℵ₀ ↔ {x | p x}.finite := lt_aleph_0_iff_set_finite From 9c444b9f6ce108032dd106601048acf279a537de Mon Sep 17 00:00:00 2001 From: Alex J Best Date: Mon, 10 Apr 2023 04:48:10 +0000 Subject: [PATCH 0801/1291] feat(.github/workflows): add modifies ported label automatically (#18782) --- .github/workflows/add_ported_warnings.yml | 25 +++++++++++++++++++++++ scripts/detect_ported_files.py | 5 +++++ 2 files changed, 30 insertions(+) diff --git a/.github/workflows/add_ported_warnings.yml b/.github/workflows/add_ported_warnings.yml index a01bef3721670..7cb0e8e22d266 100644 --- a/.github/workflows/add_ported_warnings.yml +++ b/.github/workflows/add_ported_warnings.yml @@ -25,5 +25,30 @@ jobs: uses: Ana06/get-changed-files@v1.2 - name: run the script + id: script run: | python scripts/detect_ported_files.py ${{ steps.changed-files.outputs.all }} + + - id: PR + uses: 8BitJonny/gh-get-current-pr@2.2.0 + # TODO: this may not work properly if the same commit is pushed to multiple branches: + # https://github.com/8BitJonny/gh-get-current-pr/issues/8 + with: + github-token: ${{ secrets.GITHUB_TOKEN }} + sha: ${{ github.event.pull_request.head.sha }} + # Only return if PR is still open + filterOutClosed: true + + - if: steps.script.outputs.modifies_ported == 'True' + id: add_label + name: add "modifies-synchronized-file" + # we use curl rather than octokit/request-action so that the job won't fail + # (and send an annoying email) if the labels don't exist + run: | + curl -L \ + -X POST \ + -H "Accept: application/vnd.github+json" \ + -H "Authorization: Bearer ${{ secrets.GITHUB_TOKEN }}"\ + -H "X-GitHub-Api-Version: 2022-11-28" \ + https://api.github.com/repos/${{ github.repository }}/issues/${{ steps.PR.outputs.number }}/labels \ + -d '{"labels":["modified-synchronized-file"]}' diff --git a/scripts/detect_ported_files.py b/scripts/detect_ported_files.py index b76308aa5bf31..869ad027e7b9f 100644 --- a/scripts/detect_ported_files.py +++ b/scripts/detect_ported_files.py @@ -1,5 +1,6 @@ # this script is only intended to be run by CI import sys +import os from pathlib import Path from mathlibtools.file_status import PortStatus, FileStatus @@ -18,11 +19,15 @@ def fname_for(import_path: str) -> Path: if __name__ == '__main__': files = [Path(f) for f in sys.argv[1:]] + modifies_ported = False for iname, f_status in status.file_statuses.items(): if f_status.ported: fname = fname_for(iname) if fname in files: + modifies_ported = True msg = ("Changes to this file will need to be ported to mathlib 4!\n" "Please consider retracting the changes to this file unless you are willing " "to immediately forward-port them." ) print(f"::warning file={fname},line=1,col=1::{encode_msg_text_for_github(msg)}") + with open(os.environ['GITHUB_OUTPUT'], 'a') as fh: + print(f'modifies_ported={modifies_ported}', file=fh) From 8cab1cd8f0fcaa0995b9d98188f7c7edfd4a3983 Mon Sep 17 00:00:00 2001 From: Moritz Doll Date: Mon, 10 Apr 2023 11:35:31 +0000 Subject: [PATCH 0802/1291] feat(analysis/schwartz_space): 1-dimensional derivative (#18745) --- src/analysis/schwartz_space.lean | 50 ++++++++++++++++++++++++++++---- 1 file changed, 45 insertions(+), 5 deletions(-) diff --git a/src/analysis/schwartz_space.lean b/src/analysis/schwartz_space.lean index 61f5f4ace9708..f9ac9cea0cd27 100644 --- a/src/analysis/schwartz_space.lean +++ b/src/analysis/schwartz_space.lean @@ -5,6 +5,7 @@ Authors: Moritz Doll -/ import analysis.calculus.cont_diff +import analysis.calculus.iterated_deriv import analysis.locally_convex.with_seminorms import topology.algebra.uniform_filter_basis import topology.continuous_function.bounded @@ -33,6 +34,8 @@ decay faster than any power of `‖x‖`. * `schwartz_map.seminorm`: The family of seminorms as described above * `schwartz_map.fderiv_clm`: The differential as a continuous linear map `𝓢(E, F) →L[𝕜] 𝓢(E, E →L[ℝ] F)` +* `schwartz_map.deriv_clm`: The one-dimensional derivative as a continuous linear map +`𝓢(ℝ, F) →L[𝕜] 𝓢(ℝ, F)` ## Main statements @@ -101,6 +104,10 @@ lemma smooth (f : 𝓢(E, F)) (n : ℕ∞) : cont_diff ℝ n f := f.smooth'.of_l @[protected] lemma differentiable (f : 𝓢(E, F)) : differentiable ℝ f := (f.smooth 1).differentiable rfl.le +/-- Every Schwartz function is differentiable at any point. -/ +@[protected] lemma differentiable_at (f : 𝓢(E, F)) {x : E} : differentiable_at ℝ f x := +f.differentiable.differentiable_at + @[ext] lemma ext {f g : 𝓢(E, F)} (h : ∀ x, (f : E → F) x = g x) : f = g := fun_like.ext f g h section is_O @@ -377,11 +384,31 @@ lemma seminorm_le_bound (k n : ℕ) (f : 𝓢(E, F)) {M : ℝ} (hMp: 0 ≤ M) (hM : ∀ x, ‖x‖^k * ‖iterated_fderiv ℝ n f x‖ ≤ M) : seminorm 𝕜 k n f ≤ M := f.seminorm_aux_le_bound k n hMp hM +/-- If one controls the seminorm for every `x`, then one controls the seminorm. + +Variant for functions `𝓢(ℝ, F)`. -/ +lemma seminorm_le_bound' (k n : ℕ) (f : 𝓢(ℝ, F)) {M : ℝ} (hMp: 0 ≤ M) + (hM : ∀ x, |x|^k * ‖iterated_deriv n f x‖ ≤ M) : seminorm 𝕜 k n f ≤ M := +begin + refine seminorm_le_bound 𝕜 k n f hMp _, + simpa only [real.norm_eq_abs, norm_iterated_fderiv_eq_norm_iterated_deriv], +end + /-- The seminorm controls the Schwartz estimate for any fixed `x`. -/ lemma le_seminorm (k n : ℕ) (f : 𝓢(E, F)) (x : E) : ‖x‖ ^ k * ‖iterated_fderiv ℝ n f x‖ ≤ seminorm 𝕜 k n f := f.le_seminorm_aux k n x +/-- The seminorm controls the Schwartz estimate for any fixed `x`. + +Variant for functions `𝓢(ℝ, F)`. -/ +lemma le_seminorm' (k n : ℕ) (f : 𝓢(ℝ, F)) (x : ℝ) : + |x| ^ k * ‖iterated_deriv n f x‖ ≤ seminorm 𝕜 k n f := +begin + have := le_seminorm 𝕜 k n f x, + rwa [← real.norm_eq_abs, ← norm_iterated_fderiv_eq_norm_iterated_deriv], +end + lemma norm_iterated_fderiv_le_seminorm (f : 𝓢(E, F)) (n : ℕ) (x₀ : E) : ‖iterated_fderiv ℝ n f x₀‖ ≤ (schwartz_map.seminorm 𝕜 0 n) f := begin @@ -515,18 +542,18 @@ def mk_clm [ring_hom_isometric σ] (A : (D → E) → (F → G)) end clm -section fderiv +section derivatives /-! ### Derivatives of Schwartz functions -/ variables (𝕜) variables [is_R_or_C 𝕜] [normed_space 𝕜 F] [smul_comm_class ℝ 𝕜 F] -/-- The real derivative on Schwartz space as a continuous `𝕜`-linear map. -/ +/-- The Fréchet derivative on Schwartz space as a continuous `𝕜`-linear map. -/ def fderiv_clm : 𝓢(E, F) →L[𝕜] 𝓢(E, E →L[ℝ] F) := mk_clm (fderiv ℝ) - (λ f g _, fderiv_add f.differentiable.differentiable_at g.differentiable.differentiable_at) - (λ a f _, fderiv_const_smul f.differentiable.differentiable_at a) + (λ f g _, fderiv_add f.differentiable_at g.differentiable_at) + (λ a f _, fderiv_const_smul f.differentiable_at a) (λ f, (cont_diff_top_iff_fderiv.mp f.smooth').2) (λ ⟨k, n⟩, ⟨{⟨k, n+1⟩}, 1, zero_le_one, λ f x, by simpa only [schwartz_seminorm_family_apply, seminorm.comp_apply, finset.sup_singleton, one_smul, norm_iterated_fderiv_fderiv, one_mul] @@ -535,7 +562,20 @@ mk_clm (fderiv ℝ) @[simp] lemma fderiv_clm_apply (f : 𝓢(E, F)) (x : E) : fderiv_clm 𝕜 f x = fderiv ℝ f x := rfl -end fderiv +/-- The 1-dimensional derivative on Schwartz space as a continuous `𝕜`-linear map. -/ +def deriv_clm : 𝓢(ℝ, F) →L[𝕜] 𝓢(ℝ, F) := +mk_clm (λ f, deriv f) + (λ f g _, deriv_add f.differentiable_at g.differentiable_at) + (λ a f _, deriv_const_smul a f.differentiable_at) + (λ f, (cont_diff_top_iff_deriv.mp f.smooth').2) + (λ ⟨k, n⟩, ⟨{⟨k, n+1⟩}, 1, zero_le_one, λ f x, by simpa only [real.norm_eq_abs, + finset.sup_singleton, schwartz_seminorm_family_apply, one_mul, + norm_iterated_fderiv_eq_norm_iterated_deriv, ← iterated_deriv_succ'] + using f.le_seminorm' 𝕜 k (n + 1) x⟩) + +@[simp] lemma deriv_clm_apply (f : 𝓢(ℝ, F)) (x : ℝ) : deriv_clm 𝕜 f x = deriv f x := rfl + +end derivatives section bounded_continuous_function From 0a0b3b4148b35efb8b8a38118517c5a1d30d0e69 Mon Sep 17 00:00:00 2001 From: Moritz Doll Date: Mon, 10 Apr 2023 12:40:58 +0000 Subject: [PATCH 0803/1291] feat(analysis/calculus/cont_diff): smoothness from Taylor series (#18768) The proofs are basically trivial, but I think it is very nice to have the API. --- src/analysis/calculus/cont_diff_def.lean | 14 +++++++++++++- 1 file changed, 13 insertions(+), 1 deletion(-) diff --git a/src/analysis/calculus/cont_diff_def.lean b/src/analysis/calculus/cont_diff_def.lean index 2dc7148f9e37b..b8e6ba95af231 100644 --- a/src/analysis/calculus/cont_diff_def.lean +++ b/src/analysis/calculus/cont_diff_def.lean @@ -621,6 +621,15 @@ def cont_diff_on (n : ℕ∞) (f : E → F) (s : set E) : Prop := variable {𝕜} +lemma has_ftaylor_series_up_to_on.cont_diff_on {f' : E → formal_multilinear_series 𝕜 E F} + (hf : has_ftaylor_series_up_to_on n f f' s) : cont_diff_on 𝕜 n f s := +begin + intros x hx m hm, + use s, + simp only [set.insert_eq_of_mem hx, self_mem_nhds_within, true_and], + exact ⟨f', hf.of_le hm⟩, +end + lemma cont_diff_on.cont_diff_within_at (h : cont_diff_on 𝕜 n f s) (hx : x ∈ s) : cont_diff_within_at 𝕜 n f s x := h x hx @@ -1326,6 +1335,10 @@ def cont_diff (n : ℕ∞) (f : E → F) : Prop := variable {𝕜} +/-- If `f` has a Taylor series up to `n`, then it is `C^n`. -/ +lemma has_ftaylor_series_up_to.cont_diff {f' : E → formal_multilinear_series 𝕜 E F} + (hf : has_ftaylor_series_up_to n f f') : cont_diff 𝕜 n f := ⟨f', hf⟩ + theorem cont_diff_on_univ : cont_diff_on 𝕜 n f univ ↔ cont_diff 𝕜 n f := begin split, @@ -1390,7 +1403,6 @@ lemma cont_diff_iff_forall_nat_le : cont_diff 𝕜 n f ↔ ∀ m : ℕ, ↑m ≤ n → cont_diff 𝕜 m f := by { simp_rw [← cont_diff_on_univ], exact cont_diff_on_iff_forall_nat_le } - /-! ### Iterated derivative -/ variable (𝕜) From df5e9937a06fdd349fc60106f54b84d47b1434f0 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 10 Apr 2023 12:40:59 +0000 Subject: [PATCH 0804/1291] chore(algebra/group_with_zero/units/basic): Deduplicate instance (#18785) `group_with_zero.cancel_monoid_with_zero` was a duplicate of `group_with_zero.to_cancel_monoid_with_zero`. `comm_group_with_zero.cancel_comm_monoid_with_zero` is renamed to include the usual `to_` prefix. This should have been done in #18698. --- src/algebra/group_with_zero/units/basic.lean | 12 ++---------- 1 file changed, 2 insertions(+), 10 deletions(-) diff --git a/src/algebra/group_with_zero/units/basic.lean b/src/algebra/group_with_zero/units/basic.lean index cc892e9487604..a306052d3e7c8 100644 --- a/src/algebra/group_with_zero/units/basic.lean +++ b/src/algebra/group_with_zero/units/basic.lean @@ -207,14 +207,6 @@ instance group_with_zero.no_zero_divisors : no_zero_divisors G₀ := end, .. (‹_› : group_with_zero G₀) } -@[priority 10] -- see Note [lower instance priority] -instance group_with_zero.cancel_monoid_with_zero : cancel_monoid_with_zero G₀ := -{ mul_left_cancel_of_ne_zero := λ x y z hx h, - by rw [← inv_mul_cancel_left₀ hx y, h, inv_mul_cancel_left₀ hx z], - mul_right_cancel_of_ne_zero := λ x y z hy h, - by rw [← mul_inv_cancel_right₀ hy x, h, mul_inv_cancel_right₀ hy z], - .. (‹_› : group_with_zero G₀) } - -- Can't be put next to the other `mk0` lemmas because it depends on the -- `no_zero_divisors` instance, which depends on `mk0`. @[simp] lemma units.mk0_mul (x y : G₀) (hxy) : @@ -247,8 +239,8 @@ section comm_group_with_zero -- comm variables [comm_group_with_zero G₀] {a b c d : G₀} @[priority 10] -- see Note [lower instance priority] -instance comm_group_with_zero.cancel_comm_monoid_with_zero : cancel_comm_monoid_with_zero G₀ := -{ ..group_with_zero.cancel_monoid_with_zero, ..comm_group_with_zero.to_comm_monoid_with_zero G₀ } +instance comm_group_with_zero.to_cancel_comm_monoid_with_zero : cancel_comm_monoid_with_zero G₀ := +{ ..group_with_zero.to_cancel_monoid_with_zero, ..comm_group_with_zero.to_comm_monoid_with_zero G₀ } @[priority 100] -- See note [lower instance priority] instance comm_group_with_zero.to_division_comm_monoid : division_comm_monoid G₀ := From a8c97ed34c07fcfd7ebc6b83179b8f687275eba9 Mon Sep 17 00:00:00 2001 From: sgouezel Date: Mon, 10 Apr 2023 15:40:23 +0000 Subject: [PATCH 0805/1291] feat(measure_theory/function/continuous_map_dense): compactly supported functions are dense in L^p (#18710) We have already the fact that bounded continuous functions are dense in L^p. We refactor the proof to extract an approach that is common to bounded continuous functions and to compactly supported functions, to avoid code duplication as much as possible. We also give elementary versions of the statements (in the form: for all epsilon > 0, there exists g such that...) as they may be easier to use, and specialized versions for integrable functions. Co-authored-by: sgouezel --- src/analysis/fourier/add_circle.lean | 2 +- .../function/continuous_map_dense.lean | 423 +++++++++++++----- src/measure_theory/function/lp_space.lean | 47 ++ .../function/simple_func_dense_lp.lean | 54 +++ 4 files changed, 415 insertions(+), 111 deletions(-) diff --git a/src/analysis/fourier/add_circle.lean b/src/analysis/fourier/add_circle.lean index 944f5424ac693..da480b55ef038 100644 --- a/src/analysis/fourier/add_circle.lean +++ b/src/analysis/fourier/add_circle.lean @@ -263,7 +263,7 @@ lemma coe_fn_fourier_Lp (p : ℝ≥0∞) [fact (1 ≤ p)] (n : ℤ) : lemma span_fourier_Lp_closure_eq_top {p : ℝ≥0∞} [fact (1 ≤ p)] (hp : p ≠ ∞) : (span ℂ (range (@fourier_Lp T _ p _))).topological_closure = ⊤ := begin - convert (continuous_map.to_Lp_dense_range ℂ hp (@haar_add_circle T hT) ℂ + convert (continuous_map.to_Lp_dense_range ℂ (@haar_add_circle T hT) hp ℂ ).topological_closure_map_submodule (span_fourier_closure_eq_top), rw [map_span, range_comp], simp only [continuous_linear_map.coe_coe], diff --git a/src/measure_theory/function/continuous_map_dense.lean b/src/measure_theory/function/continuous_map_dense.lean index de100258d3fd3..8e8820280f4cc 100644 --- a/src/measure_theory/function/continuous_map_dense.lean +++ b/src/measure_theory/function/continuous_map_dense.lean @@ -7,14 +7,30 @@ Authors: Heather Macbeth import measure_theory.measure.regular import measure_theory.function.simple_func_dense_lp import topology.urysohns_lemma +import measure_theory.integral.bochner /-! # Approximation in Lᵖ by continuous functions This file proves that bounded continuous functions are dense in `Lp E p μ`, for `1 ≤ p < ∞`, if the domain `α` of the functions is a normal topological space and the measure `μ` is weakly regular. +It also proves the same results for approximation by continuous functions with compact support +when the space is locally compact and `μ` is regular. + +The result is presented in several versions. First concrete versions giving an approximation +up to `ε` in these various contexts, and then abstract versions stating that the topological +closure of the relevant subgroups of `Lp` are the whole space. + +* `mem_ℒp.exists_has_compact_support_snorm_sub_le` states that, in a locally compact space, + an `ℒp` function can be approximated by continuous functions with compact support, + in the sense that `snorm (f - g) p μ` is small. +* `mem_ℒp.exists_has_compact_support_integral_rpow_sub_le`: same result, but expressed in + terms of `∫ ‖f - g‖^p`. + +Versions with `integrable` instead of `mem_ℒp` are specialized to the case `p = 1`. +Versions with `bounded_continuous` instead of `has_compact_support` drop the locally +compact assumption and give only approximation by a bounded continuous function. -The result is presented in several versions: * `measure_theory.Lp.bounded_continuous_function_dense`: The subgroup `measure_theory.Lp.bounded_continuous_function` of `Lp E p μ`, the additive subgroup of `Lp E p μ` consisting of equivalence classes containing a continuous representative, is dense in @@ -42,137 +58,324 @@ Vitali-Carathéodory theorem, in the file `measure_theory.vitali_caratheodory`. -/ open_locale ennreal nnreal topology bounded_continuous_function -open measure_theory topological_space continuous_map +open measure_theory topological_space continuous_map set variables {α : Type*} [measurable_space α] [topological_space α] [normal_space α] [borel_space α] -variables (E : Type*) [normed_add_comm_group E] - [second_countable_topology_either α E] -variables {p : ℝ≥0∞} [_i : fact (1 ≤ p)] (hp : p ≠ ∞) (μ : measure α) +variables {E : Type*} [normed_add_comm_group E] {μ : measure α} {p : ℝ≥0∞} -include _i hp - -namespace measure_theory.Lp +namespace measure_theory variables [normed_space ℝ E] +/-- A variant of Urysohn's lemma, `ℒ^p` version, for an outer regular measure `μ`: +consider two sets `s ⊆ u` which are respectively closed and open with `μ s < ∞`, and a vector `c`. +Then one may find a continuous function `f` equal to `c` on `s` and to `0` outside of `u`, +bounded by `‖c‖` everywhere, and such that the `ℒ^p` norm of `f - s.indicator (λ y, c)` is +arbitrarily small. Additionally, this function `f` belongs to `ℒ^p`. -/ +lemma exists_continuous_snorm_sub_le_of_closed [μ.outer_regular] + (hp : p ≠ ∞) {s u : set α} (s_closed : is_closed s) (u_open : is_open u) (hsu : s ⊆ u) + (hs : μ s ≠ ∞) (c : E) {ε : ℝ≥0∞} (hε : ε ≠ 0) : + ∃ (f : α → E), continuous f ∧ snorm (λ x, f x - s.indicator (λ y, c) x) p μ ≤ ε ∧ + (∀ x, ‖f x‖ ≤ ‖c‖) ∧ function.support f ⊆ u ∧ mem_ℒp f p μ := +begin + obtain ⟨η, η_pos, hη⟩ : ∃ (η : ℝ≥0), 0 < η ∧ ∀ (s : set α), μ s ≤ η → + snorm (s.indicator (λ x, c)) p μ ≤ ε, from exists_snorm_indicator_le hp c hε, + have ηpos : (0 : ℝ≥0∞) < η := ennreal.coe_lt_coe.2 η_pos, + obtain ⟨V, sV, V_open, h'V, hV⟩ : ∃ (V : set α) (H : V ⊇ s), is_open V ∧ μ V < ∞ ∧ μ (V \ s) < η, + from s_closed.measurable_set.exists_is_open_diff_lt hs ηpos.ne', + let v := u ∩ V, + have hsv : s ⊆ v := subset_inter hsu sV, + have hμv : μ v < ∞ := (measure_mono (inter_subset_right _ _)).trans_lt h'V, + obtain ⟨g, hgv, hgs, hg_range⟩ := exists_continuous_zero_one_of_closed + (u_open.inter V_open).is_closed_compl s_closed (disjoint_compl_left_iff.2 hsv), + -- Multiply this by `c` to get a continuous approximation to the function `f`; the key point is + -- that this is pointwise bounded by the indicator of the set `v \ s`, which has small measure. + have g_norm : ∀ x, ‖g x‖ = g x := λ x, by rw [real.norm_eq_abs, abs_of_nonneg (hg_range x).1], + have gc_bd0 : ∀ x, ‖g x • c‖ ≤ ‖c‖, + { assume x, + simp only [norm_smul, g_norm x], + apply mul_le_of_le_one_left (norm_nonneg _), + exact (hg_range x).2 }, + have gc_bd : ∀ x, ‖g x • c - s.indicator (λ x, c) x‖ ≤ ‖(v \ s).indicator (λ x, c) x‖, + { intros x, + by_cases hv : x ∈ v, + { rw ← set.diff_union_of_subset hsv at hv, + cases hv with hsv hs, + { simpa only [hsv.2, set.indicator_of_not_mem, not_false_iff, sub_zero, hsv, + set.indicator_of_mem] using gc_bd0 x}, + { simp [hgs hs, hs] } }, + { simp [hgv hv, (λ h, hv (hsv h) : x ∉ s)], } }, + have gc_support : function.support (λ (x : α), g x • c) ⊆ v, + { refine function.support_subset_iff'.2 (λ x hx, _), + simp only [hgv hx, pi.zero_apply, zero_smul] }, + have gc_mem : mem_ℒp (λ x, g x • c) p μ, + { apply mem_ℒp.smul_of_top_left (mem_ℒp_top_const _), + refine ⟨g.continuous.ae_strongly_measurable, _⟩, + have : snorm (v.indicator (λ x, (1 : ℝ))) p μ < ⊤, + { refine (snorm_indicator_const_le _ _).trans_lt _, + simp only [lt_top_iff_ne_top, hμv.ne, nnnorm_one, ennreal.coe_one, one_div, one_mul, ne.def, + ennreal.rpow_eq_top_iff, inv_lt_zero, false_and, or_false, not_and, not_lt, + ennreal.to_real_nonneg, implies_true_iff] }, + refine (snorm_mono (λ x, _)).trans_lt this, + by_cases hx : x ∈ v, + { simp only [hx, abs_of_nonneg (hg_range x).1, (hg_range x).2, real.norm_eq_abs, + indicator_of_mem, cstar_ring.norm_one] }, + { simp only [hgv hx, pi.zero_apply, real.norm_eq_abs, abs_zero, abs_nonneg] } }, + refine ⟨λ x, g x • c, g.continuous.smul continuous_const, (snorm_mono gc_bd).trans _, gc_bd0, + gc_support.trans (inter_subset_left _ _), gc_mem⟩, + exact hη _ ((measure_mono (diff_subset_diff (inter_subset_right _ _) subset.rfl)).trans hV.le), +end + +/-- In a locally compact space, any function in `ℒp` can be approximated by compactly supported +continuous functions when `1 ≤ p < ∞`, version in terms of `snorm`. -/ +lemma mem_ℒp.exists_has_compact_support_snorm_sub_le + [locally_compact_space α] [μ.regular] (hp : p ≠ ∞) (h'p : 1 ≤ p) + {f : α → E} (hf : mem_ℒp f p μ) {ε : ℝ≥0∞} (hε : ε ≠ 0) : + ∃ (g : α → E), has_compact_support g ∧ snorm (f - g) p μ ≤ ε ∧ continuous g ∧ mem_ℒp g p μ := +begin + suffices H : ∃ (g : α → E), snorm (f - g) p μ ≤ ε ∧ continuous g ∧ mem_ℒp g p μ ∧ + has_compact_support g, + { rcases H with ⟨g, hg, g_cont, g_mem, g_support⟩, + exact ⟨g, g_support, hg, g_cont, g_mem⟩ }, + -- It suffices to check that the set of functions we consider approximates characteristic + -- functions, is stable under addition and consists of ae strongly measurable functions. + -- First check the latter easy facts. + apply hf.induction_dense hp h'p _ _ _ _ hε, rotate, + -- stability under addition + { rintros f g ⟨f_cont, f_mem, hf⟩ ⟨g_cont, g_mem, hg⟩, + exact ⟨f_cont.add g_cont, f_mem.add g_mem, hf.add hg⟩ }, + -- ae strong measurability + { rintros f ⟨f_cont, f_mem, hf⟩, + exact f_mem.ae_strongly_measurable }, + -- We are left with approximating characteristic functions. + -- This follows from `exists_continuous_snorm_sub_le_of_closed`. + assume c t ht htμ ε hε, + have h'ε : ε / 2 ≠ 0, by simpa using hε, + obtain ⟨η, ηpos, hη⟩ : ∃ (η : ℝ≥0), 0 < η ∧ ∀ (s : set α), μ s ≤ η → + snorm (s.indicator (λ x, c)) p μ ≤ ε / 2, from exists_snorm_indicator_le hp c h'ε, + have hη_pos' : (0 : ℝ≥0∞) < η, from ennreal.coe_pos.2 ηpos, + obtain ⟨s, st, s_compact, μs⟩ : ∃ s ⊆ t, is_compact s ∧ μ (t \ s) < η, + from ht.exists_is_compact_diff_lt htμ.ne hη_pos'.ne', + have hsμ : μ s < ∞, from (measure_mono st).trans_lt htμ, + have I1 : snorm (s.indicator (λ y, c) - t.indicator (λ y, c)) p μ ≤ ε/2, + { rw [← snorm_neg, neg_sub, ← indicator_diff st], + exact (hη _ μs.le) }, + obtain ⟨k, k_compact, sk, -⟩ : ∃ (k : set α), is_compact k ∧ s ⊆ interior k ∧ k ⊆ univ, + from exists_compact_between s_compact is_open_univ (subset_univ _), + rcases exists_continuous_snorm_sub_le_of_closed hp s_compact.is_closed is_open_interior sk + hsμ.ne c h'ε with ⟨f, f_cont, I2, f_bound, f_support, f_mem⟩, + have I3 : snorm (f - t.indicator (λ y, c)) p μ ≤ ε, from calc + snorm (f - t.indicator (λ y, c)) p μ + = snorm ((f - s.indicator (λ y, c)) + (s.indicator (λ y, c) - t.indicator (λ y, c))) p μ : + by simp only [sub_add_sub_cancel] + ... ≤ snorm (f - s.indicator (λ y, c)) p μ + + snorm (s.indicator (λ y, c) - t.indicator (λ y, c)) p μ : + begin + refine snorm_add_le _ _ h'p, + { exact f_mem.ae_strongly_measurable.sub + (ae_strongly_measurable_const.indicator s_compact.measurable_set) }, + { exact (ae_strongly_measurable_const.indicator s_compact.measurable_set).sub + (ae_strongly_measurable_const.indicator ht) }, + end + ... ≤ ε/2 + ε/2 : add_le_add I2 I1 + ... = ε : ennreal.add_halves _, + refine ⟨f, I3, f_cont, f_mem, has_compact_support.intro k_compact (λ x hx, _)⟩, + rw ← function.nmem_support, + contrapose! hx, + exact interior_subset (f_support hx) +end + +/-- In a locally compact space, any function in `ℒp` can be approximated by compactly supported +continuous functions when `1 ≤ p < ∞`, version in terms of `∫`. -/ +lemma mem_ℒp.exists_has_compact_support_integral_rpow_sub_le + [locally_compact_space α] [μ.regular] {p : ℝ} (h'p : 1 ≤ p) + {f : α → E} (hf : mem_ℒp f (ennreal.of_real p) μ) {ε : ℝ} (hε : 0 < ε) : + ∃ (g : α → E), has_compact_support g ∧ ∫ x, ‖f x - g x‖^p ∂μ ≤ ε ∧ continuous g + ∧ mem_ℒp g (ennreal.of_real p) μ := +begin + have I : 0 < ε ^ (1/p) := real.rpow_pos_of_pos hε _, + have A : ennreal.of_real (ε ^ (1/p)) ≠ 0, + by simp only [ne.def, ennreal.of_real_eq_zero, not_le, I], + have B : 1 ≤ ennreal.of_real p, + { convert ennreal.of_real_le_of_real h'p, exact ennreal.of_real_one.symm }, + rcases hf.exists_has_compact_support_snorm_sub_le ennreal.coe_ne_top B A + with ⟨g, g_support, hg, g_cont, g_mem⟩, + change snorm _ (ennreal.of_real p) _ ≤ _ at hg, + refine ⟨g, g_support, _, g_cont, g_mem⟩, + rwa [(hf.sub g_mem).snorm_eq_integral_rpow_norm (zero_lt_one.trans_le B).ne' + ennreal.coe_ne_top, ennreal.of_real_le_of_real_iff I.le, one_div, + ennreal.to_real_of_real (zero_le_one.trans h'p), real.rpow_le_rpow_iff _ hε.le _] at hg, + { exact integral_nonneg (λ x, real.rpow_nonneg_of_nonneg (norm_nonneg _) _) }, + { exact inv_pos.2 (zero_lt_one.trans_le h'p) } +end + +/-- In a locally compact space, any integrable function can be approximated by compactly supported +continuous functions, version in terms of `∫⁻`. -/ +lemma integrable.exists_has_compact_support_lintegral_sub_le [locally_compact_space α] [μ.regular] + {f : α → E} (hf : integrable f μ) {ε : ℝ≥0∞} (hε : ε ≠ 0) : + ∃ (g : α → E), has_compact_support g ∧ ∫⁻ x, ‖f x - g x‖₊ ∂μ ≤ ε ∧ continuous g + ∧ integrable g μ := +begin + simp only [← mem_ℒp_one_iff_integrable, ← snorm_one_eq_lintegral_nnnorm] at hf ⊢, + exact hf.exists_has_compact_support_snorm_sub_le ennreal.one_ne_top le_rfl hε, +end + +/-- In a locally compact space, any integrable function can be approximated by compactly supported +continuous functions, version in terms of `∫`. -/ +lemma integrable.exists_has_compact_support_integral_sub_le [locally_compact_space α] [μ.regular] + {f : α → E} (hf : integrable f μ) {ε : ℝ} (hε : 0 < ε) : + ∃ (g : α → E), has_compact_support g ∧ ∫ x, ‖f x - g x‖ ∂μ ≤ ε ∧ continuous g + ∧ integrable g μ := +begin + simp only [← mem_ℒp_one_iff_integrable, ← snorm_one_eq_lintegral_nnnorm, + ← ennreal.of_real_one] at hf ⊢, + simpa using hf.exists_has_compact_support_integral_rpow_sub_le le_rfl hε, +end + +/-- Any function in `ℒp` can be approximated by bounded continuous functions when `1 ≤ p < ∞`, +version in terms of `snorm`. -/ +lemma mem_ℒp.exists_bounded_continuous_snorm_sub_le [μ.weakly_regular] (hp : p ≠ ∞) (h'p : 1 ≤ p) + {f : α → E} (hf : mem_ℒp f p μ) {ε : ℝ≥0∞} (hε : ε ≠ 0) : + ∃ (g : α →ᵇ E), snorm (f - g) p μ ≤ ε ∧ mem_ℒp g p μ := +begin + suffices H : ∃ (g : α → E), snorm (f - g) p μ ≤ ε ∧ continuous g ∧ mem_ℒp g p μ ∧ + metric.bounded (range g), + { rcases H with ⟨g, hg, g_cont, g_mem, g_bd⟩, + exact ⟨⟨⟨g, g_cont⟩, metric.bounded_range_iff.1 g_bd⟩, hg, g_mem⟩ }, + -- It suffices to check that the set of functions we consider approximates characteristic + -- functions, is stable under addition and made of ae strongly measurable functions. + -- First check the latter easy facts. + apply hf.induction_dense hp h'p _ _ _ _ hε, rotate, + -- stability under addition + { rintros f g ⟨f_cont, f_mem, f_bd⟩ ⟨g_cont, g_mem, g_bd⟩, + refine ⟨f_cont.add g_cont, f_mem.add g_mem, _⟩, + let f' : α →ᵇ E := ⟨⟨f, f_cont⟩, metric.bounded_range_iff.1 f_bd⟩, + let g' : α →ᵇ E := ⟨⟨g, g_cont⟩, metric.bounded_range_iff.1 g_bd⟩, + exact (f' + g').bounded_range }, + -- ae strong measurability + { exact λ f ⟨_, h, _⟩, h.ae_strongly_measurable }, + -- We are left with approximating characteristic functions. + -- This follows from `exists_continuous_snorm_sub_le_of_closed`. + assume c t ht htμ ε hε, + have h'ε : ε / 2 ≠ 0, by simpa using hε, + obtain ⟨η, ηpos, hη⟩ : ∃ (η : ℝ≥0), 0 < η ∧ ∀ (s : set α), μ s ≤ η → + snorm (s.indicator (λ x, c)) p μ ≤ ε / 2, from exists_snorm_indicator_le hp c h'ε, + have hη_pos' : (0 : ℝ≥0∞) < η, from ennreal.coe_pos.2 ηpos, + obtain ⟨s, st, s_closed, μs⟩ : ∃ s ⊆ t, is_closed s ∧ μ (t \ s) < η, + from ht.exists_is_closed_diff_lt htμ.ne hη_pos'.ne', + have hsμ : μ s < ∞, from (measure_mono st).trans_lt htμ, + have I1 : snorm (s.indicator (λ y, c) - t.indicator (λ y, c)) p μ ≤ ε/2, + { rw [← snorm_neg, neg_sub, ← indicator_diff st], + exact (hη _ μs.le) }, + rcases exists_continuous_snorm_sub_le_of_closed hp s_closed is_open_univ (subset_univ _) + hsμ.ne c h'ε with ⟨f, f_cont, I2, f_bound, -, f_mem⟩, + have I3 : snorm (f - t.indicator (λ y, c)) p μ ≤ ε, from calc + snorm (f - t.indicator (λ y, c)) p μ + = snorm ((f - s.indicator (λ y, c)) + (s.indicator (λ y, c) - t.indicator (λ y, c))) p μ : + by simp only [sub_add_sub_cancel] + ... ≤ snorm (f - s.indicator (λ y, c)) p μ + + snorm (s.indicator (λ y, c) - t.indicator (λ y, c)) p μ : + begin + refine snorm_add_le _ _ h'p, + { exact f_mem.ae_strongly_measurable.sub + (ae_strongly_measurable_const.indicator s_closed.measurable_set) }, + { exact (ae_strongly_measurable_const.indicator s_closed.measurable_set).sub + (ae_strongly_measurable_const.indicator ht) }, + end + ... ≤ ε/2 + ε/2 : add_le_add I2 I1 + ... = ε : ennreal.add_halves _, + refine ⟨f, I3, f_cont, f_mem, _⟩, + exact (bounded_continuous_function.of_normed_add_comm_group f f_cont _ f_bound).bounded_range, +end + +/-- Any function in `ℒp` can be approximated by bounded continuous functions when `1 ≤ p < ∞`, +version in terms of `∫`. -/ +lemma mem_ℒp.exists_bounded_continuous_integral_rpow_sub_le + [μ.weakly_regular] {p : ℝ} (h'p : 1 ≤ p) + {f : α → E} (hf : mem_ℒp f (ennreal.of_real p) μ) {ε : ℝ} (hε : 0 < ε) : + ∃ (g : α →ᵇ E), ∫ x, ‖f x - g x‖^p ∂μ ≤ ε ∧ mem_ℒp g (ennreal.of_real p) μ := +begin + have I : 0 < ε ^ (1/p) := real.rpow_pos_of_pos hε _, + have A : ennreal.of_real (ε ^ (1/p)) ≠ 0, + by simp only [ne.def, ennreal.of_real_eq_zero, not_le, I], + have B : 1 ≤ ennreal.of_real p, + { convert ennreal.of_real_le_of_real h'p, exact ennreal.of_real_one.symm }, + rcases hf.exists_bounded_continuous_snorm_sub_le ennreal.coe_ne_top B A + with ⟨g, hg, g_mem⟩, + change snorm _ (ennreal.of_real p) _ ≤ _ at hg, + refine ⟨g, _, g_mem⟩, + rwa [(hf.sub g_mem).snorm_eq_integral_rpow_norm (zero_lt_one.trans_le B).ne' + ennreal.coe_ne_top, ennreal.of_real_le_of_real_iff I.le, one_div, + ennreal.to_real_of_real (zero_le_one.trans h'p), real.rpow_le_rpow_iff _ hε.le _] at hg, + { exact integral_nonneg (λ x, real.rpow_nonneg_of_nonneg (norm_nonneg _) _) }, + { exact inv_pos.2 (zero_lt_one.trans_le h'p) } +end + +/-- Any integrable function can be approximated by bounded continuous functions, +version in terms of `∫⁻`. -/ +lemma integrable.exists_bounded_continuous_lintegral_sub_le [μ.weakly_regular] + {f : α → E} (hf : integrable f μ) {ε : ℝ≥0∞} (hε : ε ≠ 0) : + ∃ (g : α →ᵇ E), ∫⁻ x, ‖f x - g x‖₊ ∂μ ≤ ε ∧ integrable g μ := +begin + simp only [← mem_ℒp_one_iff_integrable, ← snorm_one_eq_lintegral_nnnorm] at hf ⊢, + exact hf.exists_bounded_continuous_snorm_sub_le ennreal.one_ne_top le_rfl hε, +end + +/-- Any integrable function can be approximated by bounded continuous functions, +version in terms of `∫`. -/ +lemma integrable.exists_bounded_continuous_integral_sub_le [μ.weakly_regular] + {f : α → E} (hf : integrable f μ) {ε : ℝ} (hε : 0 < ε) : + ∃ (g : α →ᵇ E), ∫ x, ‖f x - g x‖ ∂μ ≤ ε ∧ integrable g μ := +begin + simp only [← mem_ℒp_one_iff_integrable, ← snorm_one_eq_lintegral_nnnorm, + ← ennreal.of_real_one] at hf ⊢, + simpa using hf.exists_bounded_continuous_integral_rpow_sub_le le_rfl hε, +end + +namespace Lp + +variables (E) + /-- A function in `Lp` can be approximated in `Lp` by continuous functions. -/ -lemma bounded_continuous_function_dense [μ.weakly_regular] : +lemma bounded_continuous_function_dense + [second_countable_topology_either α E] [_i : fact (1 ≤ p)] (hp : p ≠ ∞) [μ.weakly_regular] : (bounded_continuous_function E p μ).topological_closure = ⊤ := begin - have hp₀ : 0 < p := lt_of_lt_of_le zero_lt_one _i.elim, - have hp₀' : 0 ≤ 1 / p.to_real := div_nonneg zero_le_one ennreal.to_real_nonneg, - have hp₀'' : 0 < p.to_real, - { simpa [← ennreal.to_real_lt_to_real ennreal.zero_ne_top hp] using hp₀ }, - -- It suffices to prove that scalar multiples of the indicator function of a finite-measure - -- measurable set can be approximated by continuous functions - suffices : ∀ (c : E) {s : set α} (hs : measurable_set s) (hμs : μ s < ⊤), - (Lp.simple_func.indicator_const p hs hμs.ne c : Lp E p μ) - ∈ (bounded_continuous_function E p μ).topological_closure, - { rw add_subgroup.eq_top_iff', - refine Lp.induction hp _ _ _ _, - { exact this }, - { exact λ f g hf hg hfg', add_subgroup.add_mem _ }, - { exact add_subgroup.is_closed_topological_closure _ } }, - -- Let `s` be a finite-measure measurable set, let's approximate `c` times its indicator function - intros c s hs hsμ, + rw add_subgroup.eq_top_iff', + assume f, refine mem_closure_iff_frequently.mpr _, rw metric.nhds_basis_closed_ball.frequently_iff, intros ε hε, - -- A little bit of pre-emptive work, to find `η : ℝ≥0` which will be a margin small enough for - -- our purposes - obtain ⟨η, hη_pos, hη_le⟩ : ∃ η, 0 < η ∧ (↑(‖bit0 (‖c‖)‖₊ * (2 * η) ^ (1 / p.to_real)) : ℝ) ≤ ε, - { have : filter.tendsto (λ x : ℝ≥0, ‖bit0 (‖c‖)‖₊ * (2 * x) ^ (1 / p.to_real)) (𝓝 0) (𝓝 0), - { have : filter.tendsto (λ x : ℝ≥0, 2 * x) (𝓝 0) (𝓝 (2 * 0)) := filter.tendsto_id.const_mul 2, - convert ((nnreal.continuous_at_rpow_const (or.inr hp₀')).tendsto.comp this).const_mul _, - simp [hp₀''.ne'] }, - let ε' : ℝ≥0 := ⟨ε, hε.le⟩, - have hε' : 0 < ε' := by exact_mod_cast hε, - obtain ⟨δ, hδ, hδε'⟩ := - nnreal.nhds_zero_basis.eventually_iff.mp (eventually_le_of_tendsto_lt hε' this), - obtain ⟨η, hη, hηδ⟩ := exists_between hδ, - refine ⟨η, hη, _⟩, - exact_mod_cast hδε' hηδ }, - have hη_pos' : (0 : ℝ≥0∞) < η := ennreal.coe_pos.2 hη_pos, - -- Use the regularity of the measure to `η`-approximate `s` by an open superset and a closed - -- subset - obtain ⟨u, su, u_open, μu⟩ : ∃ u ⊇ s, is_open u ∧ μ u < μ s + ↑η, - { refine s.exists_is_open_lt_of_lt _ _, - simpa using ennreal.add_lt_add_left hsμ.ne hη_pos' }, - obtain ⟨F, Fs, F_closed, μF⟩ : ∃ F ⊆ s, is_closed F ∧ μ s < μ F + ↑η := - hs.exists_is_closed_lt_add hsμ.ne hη_pos'.ne', - have : disjoint uᶜ F := (Fs.trans su).disjoint_compl_left, - have h_μ_sdiff : μ (u \ F) ≤ 2 * η, - { have hFμ : μ F < ⊤ := (measure_mono Fs).trans_lt hsμ, - refine ennreal.le_of_add_le_add_left hFμ.ne _, - have : μ u < μ F + ↑η + ↑η, - from μu.trans (ennreal.add_lt_add_right ennreal.coe_ne_top μF), - convert this.le using 1, - { rw [add_comm, ← measure_union, set.diff_union_of_subset (Fs.trans su)], - exacts [disjoint_sdiff_self_left, F_closed.measurable_set] }, - have : (2:ℝ≥0∞) * η = η + η := by simpa using add_mul (1:ℝ≥0∞) 1 η, - rw this, - abel }, - -- Apply Urysohn's lemma to get a continuous approximation to the characteristic function of - -- the set `s` - obtain ⟨g, hgu, hgF, hg_range⟩ := - exists_continuous_zero_one_of_closed u_open.is_closed_compl F_closed this, - -- Multiply this by `c` to get a continuous approximation to the function `f`; the key point is - -- that this is pointwise bounded by the indicator of the set `u \ F` - have g_norm : ∀ x, ‖g x‖ = g x := λ x, by rw [real.norm_eq_abs, abs_of_nonneg (hg_range x).1], - have gc_bd : ∀ x, ‖g x • c - s.indicator (λ x, c) x‖ ≤ ‖(u \ F).indicator (λ x, bit0 ‖c‖) x‖, - { intros x, - by_cases hu : x ∈ u, - { rw ← set.diff_union_of_subset (Fs.trans su) at hu, - cases hu with hFu hF, - { refine (norm_sub_le _ _).trans _, - refine (add_le_add_left (norm_indicator_le_norm_self (λ x, c) x) _).trans _, - have h₀ : g x * ‖c‖ + ‖c‖ ≤ 2 * ‖c‖, - { nlinarith [(hg_range x).1, (hg_range x).2, norm_nonneg c] }, - have h₁ : (2:ℝ) * ‖c‖ = bit0 (‖c‖) := by simpa using add_mul (1:ℝ) 1 (‖c‖), - simp [hFu, norm_smul, h₀, ← h₁, g_norm x] }, - { simp [hgF hF, Fs hF] } }, - { have : x ∉ s := λ h, hu (su h), - simp [hgu hu, this] } }, - -- The rest is basically just `ennreal`-arithmetic - have gc_snorm : snorm ((λ x, g x • c) - s.indicator (λ x, c)) p μ - ≤ (↑(‖bit0 (‖c‖)‖₊ * (2 * η) ^ (1 / p.to_real)) : ℝ≥0∞), - { refine (snorm_mono_ae (filter.eventually_of_forall gc_bd)).trans _, - rw snorm_indicator_const (u_open.sdiff F_closed).measurable_set hp₀.ne' hp, - push_cast [← ennreal.coe_rpow_of_nonneg _ hp₀'], - exact ennreal.mul_left_mono (ennreal.monotone_rpow_of_nonneg hp₀' h_μ_sdiff) }, - have gc_cont : continuous (λ x, g x • c) := g.continuous.smul continuous_const, - have gc_mem_ℒp : mem_ℒp (λ x, g x • c) p μ, - { have : mem_ℒp ((λ x, g x • c) - s.indicator (λ x, c)) p μ := - ⟨gc_cont.ae_strongly_measurable.sub (strongly_measurable_const.indicator hs) - .ae_strongly_measurable, - gc_snorm.trans_lt ennreal.coe_lt_top⟩, - simpa using this.add (mem_ℒp_indicator_const p hs c (or.inr hsμ.ne)) }, - refine ⟨gc_mem_ℒp.to_Lp _, _, _⟩, - { rw mem_closed_ball_iff_norm, - refine le_trans _ hη_le, - rw [simple_func.coe_indicator_const, indicator_const_Lp, ← mem_ℒp.to_Lp_sub, Lp.norm_to_Lp], - exact ennreal.to_real_le_coe_of_le_coe gc_snorm }, - { rw [set_like.mem_coe, mem_bounded_continuous_function_iff], - refine ⟨bounded_continuous_function.of_normed_add_comm_group _ gc_cont (‖c‖) _, rfl⟩, - intros x, - have h₀ : g x * ‖c‖ ≤ ‖c‖, - { nlinarith [(hg_range x).1, (hg_range x).2, norm_nonneg c] }, - simp [norm_smul, g_norm x, h₀] }, + have A : ennreal.of_real ε ≠ 0, by simp only [ne.def, ennreal.of_real_eq_zero, not_le, hε], + obtain ⟨g, hg, g_mem⟩ : ∃ (g : α →ᵇ E), snorm (f - g) p μ ≤ ennreal.of_real ε ∧ mem_ℒp g p μ, + from (Lp.mem_ℒp f).exists_bounded_continuous_snorm_sub_le hp _i.out A, + refine ⟨g_mem.to_Lp _, _, ⟨g, rfl⟩⟩, + simp only [dist_eq_norm, metric.mem_closed_ball'], + rw Lp.norm_def, + convert ennreal.to_real_le_of_le_of_real hε.le hg using 2, + apply snorm_congr_ae, + filter_upwards [coe_fn_sub f (g_mem.to_Lp g), g_mem.coe_fn_to_Lp] with x hx h'x, + simp only [hx, pi.sub_apply, sub_right_inj, h'x], end -end measure_theory.Lp +end Lp +end measure_theory + +variables [second_countable_topology_either α E] [_i : fact (1 ≤ p)] (hp : p ≠ ∞) variables (𝕜 : Type*) [normed_field 𝕜] [normed_algebra ℝ 𝕜] [normed_space 𝕜 E] +include _i hp +variables (E) (μ) namespace bounded_continuous_function -open linear_map (range) lemma to_Lp_dense_range [μ.weakly_regular] [is_finite_measure μ] : dense_range ⇑(to_Lp p μ 𝕜 : (α →ᵇ E) →L[𝕜] Lp E p μ) := begin haveI : normed_space ℝ E := restrict_scalars.normed_space ℝ 𝕜 E, rw dense_range_iff_closure_range, - suffices : (range (to_Lp p μ 𝕜 : _ →L[𝕜] Lp E p μ)).to_add_subgroup.topological_closure = ⊤, + suffices : (linear_map.range (to_Lp p μ 𝕜 : _ →L[𝕜] Lp E p μ)) + .to_add_subgroup.topological_closure = ⊤, { exact congr_arg coe this }, simp [range_to_Lp p μ, measure_theory.Lp.bounded_continuous_function_dense E hp], end @@ -180,14 +383,14 @@ end end bounded_continuous_function namespace continuous_map -open linear_map (range) lemma to_Lp_dense_range [compact_space α] [μ.weakly_regular] [is_finite_measure μ] : dense_range ⇑(to_Lp p μ 𝕜 : C(α, E) →L[𝕜] Lp E p μ) := begin haveI : normed_space ℝ E := restrict_scalars.normed_space ℝ 𝕜 E, rw dense_range_iff_closure_range, - suffices : (range (to_Lp p μ 𝕜 : _ →L[𝕜] Lp E p μ)).to_add_subgroup.topological_closure = ⊤, + suffices : (linear_map.range (to_Lp p μ 𝕜 : _ →L[𝕜] Lp E p μ)) + .to_add_subgroup.topological_closure = ⊤, { exact congr_arg coe this }, simp [range_to_Lp p μ, measure_theory.Lp.bounded_continuous_function_dense E hp] end diff --git a/src/measure_theory/function/lp_space.lean b/src/measure_theory/function/lp_space.lean index b4b9c72fdc412..ab45f937e1a84 100644 --- a/src/measure_theory/function/lp_space.lean +++ b/src/measure_theory/function/lp_space.lean @@ -1874,6 +1874,23 @@ begin { exact snorm_indicator_const hs hp hp_top, }, end +lemma snorm_indicator_const_le (c : G) (p : ℝ≥0∞) : + snorm (s.indicator (λ x, c)) p μ ≤ ‖c‖₊ * (μ s) ^ (1 / p.to_real) := +begin + rcases eq_or_ne p 0 with rfl|hp, + { simp only [snorm_exponent_zero, zero_le'] }, + rcases eq_or_ne p ∞ with rfl|h'p, + { simp only [snorm_exponent_top, ennreal.top_to_real, div_zero, ennreal.rpow_zero, mul_one], + exact snorm_ess_sup_indicator_const_le _ _ }, + let t := to_measurable μ s, + calc snorm (s.indicator (λ x, c)) p μ + ≤ snorm (t.indicator (λ x, c)) p μ : + snorm_mono (norm_indicator_le_of_subset (subset_to_measurable _ _) _) + ... = ‖c‖₊ * (μ t) ^ (1 / p.to_real) : + snorm_indicator_const (measurable_set_to_measurable _ _) hp h'p + ... = ‖c‖₊ * (μ s) ^ (1 / p.to_real) : by rw measure_to_measurable +end + lemma mem_ℒp.indicator (hs : measurable_set s) (hf : mem_ℒp f p μ) : mem_ℒp (s.indicator f) p μ := ⟨hf.ae_strongly_measurable.indicator hs, lt_of_le_of_lt (snorm_indicator_le f) hf.snorm_lt_top⟩ @@ -1933,6 +1950,36 @@ begin { exact or.inr hμsc.lt_top, }, end +/-- The `ℒ^p` norm of the indicator of a set is uniformly small if the set itself has small measure, +for any `p < ∞`. Given here as an existential `∀ ε > 0, ∃ η > 0, ...` to avoid later +management of `ℝ≥0∞`-arithmetic. -/ +lemma exists_snorm_indicator_le (hp : p ≠ ∞) (c : E) {ε : ℝ≥0∞} (hε : ε ≠ 0) : + ∃ (η : ℝ≥0), 0 < η ∧ ∀ (s : set α), μ s ≤ η → snorm (s.indicator (λ x, c)) p μ ≤ ε := +begin + rcases eq_or_ne p 0 with rfl|h'p, + { exact ⟨1, zero_lt_one, λ s hs, by simp⟩ }, + have hp₀ : 0 < p := bot_lt_iff_ne_bot.2 h'p, + have hp₀' : 0 ≤ 1 / p.to_real := div_nonneg zero_le_one ennreal.to_real_nonneg, + have hp₀'' : 0 < p.to_real, + { simpa [← ennreal.to_real_lt_to_real ennreal.zero_ne_top hp] using hp₀ }, + obtain ⟨η, hη_pos, hη_le⟩ : ∃ (η : ℝ≥0), 0 < η ∧ (‖c‖₊ * η ^ (1 / p.to_real) : ℝ≥0∞) ≤ ε, + { have : filter.tendsto (λ x : ℝ≥0, ((‖c‖₊ * x ^ (1 / p.to_real) : ℝ≥0) : ℝ≥0∞)) + (𝓝 0) (𝓝 (0 : ℝ≥0)), + { rw ennreal.tendsto_coe, + convert ((nnreal.continuous_at_rpow_const (or.inr hp₀')).tendsto).const_mul _, + simp [hp₀''.ne'] }, + have hε' : 0 < ε := hε.bot_lt, + obtain ⟨δ, hδ, hδε'⟩ := + nnreal.nhds_zero_basis.eventually_iff.mp (eventually_le_of_tendsto_lt hε' this), + obtain ⟨η, hη, hηδ⟩ := exists_between hδ, + refine ⟨η, hη, _⟩, + rw [ennreal.coe_rpow_of_nonneg _ hp₀', ← ennreal.coe_mul], + exact hδε' hηδ }, + refine ⟨η, hη_pos, λ s hs, _⟩, + refine (snorm_indicator_const_le _ _).trans (le_trans _ hη_le), + exact mul_le_mul_left' (ennreal.rpow_le_rpow hs hp₀') _, +end + end indicator section indicator_const_Lp diff --git a/src/measure_theory/function/simple_func_dense_lp.lean b/src/measure_theory/function/simple_func_dense_lp.lean index 5c0400234c670..b8a92f376fba8 100644 --- a/src/measure_theory/function/simple_func_dense_lp.lean +++ b/src/measure_theory/function/simple_func_dense_lp.lean @@ -895,6 +895,60 @@ begin exact λ f hf, h_ae hf.coe_fn_to_Lp (Lp.mem_ℒp _) (this (hf.to_Lp f)), end +/-- If a set of ae strongly measurable functions is stable under addition and approximates +characteristic functions in `ℒp`, then it is dense in `ℒp`. -/ +lemma mem_ℒp.induction_dense (hp_ne_top : p ≠ ∞) (h'p : 1 ≤ p) (P : (α → E) → Prop) + (h0P : ∀ (c : E) ⦃s : set α⦄, measurable_set s → μ s < ∞ → ∀ {ε : ℝ≥0∞}, ε ≠ 0 → + (∃ (g : α → E), snorm (g - s.indicator (λ x, c)) p μ ≤ ε ∧ P g)) + (h1P : ∀ f g, P f → P g → P (f + g)) + (h2P : ∀ f, P f → ae_strongly_measurable f μ) + {f : α → E} (hf : mem_ℒp f p μ) {ε : ℝ≥0∞} (hε : ε ≠ 0) : + ∃ (g : α → E), snorm (f - g) p μ ≤ ε ∧ P g := +begin + haveI : fact (1 ≤ p) := ⟨h'p⟩, + revert f hf ε, + refine mem_ℒp.induction hp_ne_top _ _ _ _ _, + { assume c s hs hμs ε εpos, + rcases h0P c hs hμs εpos with ⟨g, hg, Pg⟩, + rw [← snorm_neg, neg_sub] at hg, + exact ⟨g, hg, Pg⟩ }, + { assume f f' hff' hf hf' Hf Hf' ε εpos, + have A : ε / 2 ≠ 0, by simp [εpos], + rcases Hf A with ⟨g, hfg, Pg⟩, + rcases Hf' A with ⟨g', hf'g', Pg'⟩, + refine ⟨g + g', _, h1P g g' Pg Pg'⟩, + calc snorm (f + f' - (g + g')) p μ + = snorm ((f - g) + (f' - g')) p μ : by { congr' 1, abel } + ... ≤ snorm (f - g) p μ + snorm (f' - g') p μ : + snorm_add_le (hf.ae_strongly_measurable.sub (h2P g Pg)) + (hf'.ae_strongly_measurable.sub (h2P g' Pg')) h'p + ... ≤ ε / 2 + ε / 2 : add_le_add hfg hf'g' + ... = ε : ennreal.add_halves _ }, + { rw is_closed_iff_nhds, + assume f hf ε εpos, + have A : ε / 2 ≠ 0, by simp [εpos], + rcases hf (emetric.ball f (ε/2)) (emetric.ball_mem_nhds _ A.bot_lt) with ⟨f', hf', h'f'⟩, + rcases h'f' A with ⟨g, hg, Pg⟩, + refine ⟨g, _, Pg⟩, + calc snorm (f - g) p μ = snorm ((f - f') + (f' - g)) p μ : by simp only [sub_add_sub_cancel] + ... ≤ snorm (f - f') p μ + snorm (f' - g) p μ : + snorm_add_le ((Lp.mem_ℒp f).sub (Lp.mem_ℒp f')).ae_strongly_measurable + ((Lp.mem_ℒp f').ae_strongly_measurable.sub (h2P g Pg)) h'p + ... ≤ ε / 2 + ε / 2 : + begin + refine add_le_add _ hg, + rw [← snorm_neg, neg_sub], + simp only [Lp.edist_def, emetric.mem_ball] at hf', + exact hf'.le + end + ... = ε : ennreal.add_halves _ }, + { assume f f' hff' hf Hf ε εpos, + rcases Hf εpos with ⟨g, hg, Pg⟩, + refine ⟨g, _, Pg⟩, + have : f - g =ᵐ[μ] f' - g := hff'.sub (filter.germ.coe_eq.mp rfl), + rwa ← snorm_congr_ae this } +end + section integrable notation α ` →₁ₛ[`:25 μ `] ` E := @measure_theory.Lp.simple_func α E _ _ 1 μ From c9236f47f5b9df573443aa499c0d3968769628b7 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 10 Apr 2023 17:00:44 +0000 Subject: [PATCH 0806/1291] ci: fix label name for ported files (#18788) --- .github/workflows/add_ported_warnings.yml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/.github/workflows/add_ported_warnings.yml b/.github/workflows/add_ported_warnings.yml index 7cb0e8e22d266..7530aa696e99c 100644 --- a/.github/workflows/add_ported_warnings.yml +++ b/.github/workflows/add_ported_warnings.yml @@ -51,4 +51,4 @@ jobs: -H "Authorization: Bearer ${{ secrets.GITHUB_TOKEN }}"\ -H "X-GitHub-Api-Version: 2022-11-28" \ https://api.github.com/repos/${{ github.repository }}/issues/${{ steps.PR.outputs.number }}/labels \ - -d '{"labels":["modified-synchronized-file"]}' + -d '{"labels":["modifies-synchronized-file"]}' From 37ffa4ee6ae02f5f6ca7226922143d3a10961e3d Mon Sep 17 00:00:00 2001 From: Christopher Hoskin Date: Mon, 10 Apr 2023 21:54:43 +0000 Subject: [PATCH 0807/1291] refactor(topology.continuous_function.algebra): Loosen continuous_subsemiring hypothesis (#18790) The current hypothesis of `continuous_subsemiring` and `continuous_submonoid` are unnecessarily restrictive. --- src/topology/continuous_function/algebra.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/topology/continuous_function/algebra.lean b/src/topology/continuous_function/algebra.lean index 29d3335e29e17..bf4e1f25a46af 100644 --- a/src/topology/continuous_function/algebra.lean +++ b/src/topology/continuous_function/algebra.lean @@ -208,7 +208,7 @@ section subtype /-- The `submonoid` of continuous maps `α → β`. -/ @[to_additive "The `add_submonoid` of continuous maps `α → β`. "] def continuous_submonoid (α : Type*) (β : Type*) [topological_space α] [topological_space β] - [monoid β] [has_continuous_mul β] : submonoid (α → β) := + [mul_one_class β] [has_continuous_mul β] : submonoid (α → β) := { carrier := { f : α → β | continuous f }, one_mem' := @continuous_const _ _ _ _ 1, mul_mem' := λ f g fc gc, fc.mul gc } @@ -379,7 +379,7 @@ section subtype /-- The subsemiring of continuous maps `α → β`. -/ def continuous_subsemiring (α : Type*) (R : Type*) [topological_space α] [topological_space R] - [semiring R] [topological_semiring R] : subsemiring (α → R) := + [non_assoc_semiring R] [topological_semiring R] : subsemiring (α → R) := { ..continuous_add_submonoid α R, ..continuous_submonoid α R } From a493616c740a3252e4cd0e4d0851984946b7b268 Mon Sep 17 00:00:00 2001 From: Moritz Doll Date: Tue, 11 Apr 2023 06:22:06 +0000 Subject: [PATCH 0808/1291] feat(analysis/calculus/cont_diff_def): C^(n+1) iff C^n derivative (#18767) --- src/analysis/calculus/cont_diff_def.lean | 17 +++++++++++++++++ 1 file changed, 17 insertions(+) diff --git a/src/analysis/calculus/cont_diff_def.lean b/src/analysis/calculus/cont_diff_def.lean index b8e6ba95af231..2d4bf3367fc70 100644 --- a/src/analysis/calculus/cont_diff_def.lean +++ b/src/analysis/calculus/cont_diff_def.lean @@ -1087,6 +1087,17 @@ begin rwa fderiv_within_inter (is_open.mem_nhds o_open hy.2) (hs y hy.1) at A end +lemma cont_diff_on_succ_iff_has_fderiv_within {n : ℕ} (hs : unique_diff_on 𝕜 s) : + cont_diff_on 𝕜 ((n + 1) : ℕ) f s ↔ ∃ (f' : E → (E →L[𝕜] F)), + cont_diff_on 𝕜 n f' s ∧ ∀ x, x ∈ s → has_fderiv_within_at f (f' x) s x := +begin + rw cont_diff_on_succ_iff_fderiv_within hs, + refine ⟨λ h, ⟨fderiv_within 𝕜 f s, h.2, λ x hx, (h.1 x hx).has_fderiv_within_at⟩, λ h, _⟩, + rcases h with ⟨f', h1, h2⟩, + refine ⟨λ x hx, (h2 x hx).differentiable_within_at, λ x hx, _⟩, + exact (h1 x hx).congr' (λ y hy, (h2 y hy).fderiv_within (hs y hy)) hx, +end + /-- A function is `C^(n + 1)` on an open domain if and only if it is differentiable there, and its derivative (expressed with `fderiv`) is `C^n`. -/ theorem cont_diff_on_succ_iff_fderiv_of_open {n : ℕ} (hs : is_open s) : @@ -1403,6 +1414,12 @@ lemma cont_diff_iff_forall_nat_le : cont_diff 𝕜 n f ↔ ∀ m : ℕ, ↑m ≤ n → cont_diff 𝕜 m f := by { simp_rw [← cont_diff_on_univ], exact cont_diff_on_iff_forall_nat_le } +/-- A function is `C^(n+1)` iff it has a `C^n` derivative. -/ +lemma cont_diff_succ_iff_has_fderiv {n : ℕ} : cont_diff 𝕜 ((n + 1) : ℕ) f ↔ + ∃ (f' : E → (E →L[𝕜] F)), cont_diff 𝕜 n f' ∧ ∀ x, has_fderiv_at f (f' x) x := +by simp only [← cont_diff_on_univ, ← has_fderiv_within_at_univ, + cont_diff_on_succ_iff_has_fderiv_within (unique_diff_on_univ), set.mem_univ, forall_true_left] + /-! ### Iterated derivative -/ variable (𝕜) From 25a9423c6b2c8626e91c688bfd6c1d0a986a3e6e Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Tue, 11 Apr 2023 07:50:43 +0000 Subject: [PATCH 0809/1291] chore(*): add mathlib4 synchronization comments (#18772) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.module.projective` * `algebra.order.rearrangement` * `analysis.convex.function` * `analysis.locally_convex.polar` * `category_theory.Fintype` * `category_theory.abelian.functor_category` * `category_theory.sites.plus` * `category_theory.triangulated.pretriangulated` * `category_theory.triangulated.rotate` * `category_theory.triangulated.triangulated` * `control.fold` * `data.finset.sups` * `data.list.to_finsupp` * `data.num.lemmas` * `field_theory.finiteness` * `field_theory.mv_polynomial` * `linear_algebra.alternating` * `linear_algebra.dimension` * `linear_algebra.free_module.finite.basic` * `order.height` * `set_theory.zfc.ordinal` * `topology.algebra.module.weak_dual` * `topology.local_at_target` * `topology.metric_space.contracting` * `topology.metric_space.pi_nat` --- src/algebra/module/projective.lean | 3 +++ src/algebra/order/rearrangement.lean | 3 +++ src/analysis/convex/function.lean | 3 +++ src/analysis/locally_convex/polar.lean | 3 +++ src/category_theory/Fintype.lean | 3 +++ src/category_theory/abelian/functor_category.lean | 3 +++ src/category_theory/sites/plus.lean | 3 +++ src/category_theory/triangulated/pretriangulated.lean | 3 +++ src/category_theory/triangulated/rotate.lean | 3 +++ src/category_theory/triangulated/triangulated.lean | 3 +++ src/control/fold.lean | 3 +++ src/data/finset/sups.lean | 3 +++ src/data/list/to_finsupp.lean | 3 +++ src/data/num/lemmas.lean | 3 +++ src/field_theory/finiteness.lean | 3 +++ src/field_theory/mv_polynomial.lean | 3 +++ src/linear_algebra/alternating.lean | 3 +++ src/linear_algebra/dimension.lean | 3 +++ src/linear_algebra/free_module/finite/basic.lean | 3 +++ src/order/height.lean | 3 +++ src/set_theory/zfc/ordinal.lean | 3 +++ src/topology/algebra/module/weak_dual.lean | 3 +++ src/topology/local_at_target.lean | 3 +++ src/topology/metric_space/contracting.lean | 3 +++ src/topology/metric_space/pi_nat.lean | 3 +++ 25 files changed, 75 insertions(+) diff --git a/src/algebra/module/projective.lean b/src/algebra/module/projective.lean index 3cc28e42cf214..c147cd4606636 100644 --- a/src/algebra/module/projective.lean +++ b/src/algebra/module/projective.lean @@ -12,6 +12,9 @@ import linear_algebra.free_module.basic # Projective modules +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains a definition of a projective module, the proof that our definition is equivalent to a lifting property, and the proof that all free modules are projective. diff --git a/src/algebra/order/rearrangement.lean b/src/algebra/order/rearrangement.lean index f4575a8d4414c..9cd2d6d035fda 100644 --- a/src/algebra/order/rearrangement.lean +++ b/src/algebra/order/rearrangement.lean @@ -13,6 +13,9 @@ import tactic.abel /-! # Rearrangement inequality +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves the rearrangement inequality and deduces the conditions for equality and strict inequality. diff --git a/src/analysis/convex/function.lean b/src/analysis/convex/function.lean index ef5948f6e871f..bd2752936998c 100644 --- a/src/analysis/convex/function.lean +++ b/src/analysis/convex/function.lean @@ -8,6 +8,9 @@ import analysis.convex.basic /-! # Convex and concave functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines convex and concave functions in vector spaces and proves the finite Jensen inequality. The integral version can be found in `analysis.convex.integral`. diff --git a/src/analysis/locally_convex/polar.lean b/src/analysis/locally_convex/polar.lean index f02fa411c5415..a52213b20468d 100644 --- a/src/analysis/locally_convex/polar.lean +++ b/src/analysis/locally_convex/polar.lean @@ -11,6 +11,9 @@ import topology.algebra.module.weak_dual /-! # Polar set +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define the polar set. There are different notions of the polar, we will define the *absolute polar*. The advantage over the real polar is that we can define the absolute polar for any bilinear form `B : E →ₗ[𝕜] F →ₗ[𝕜] 𝕜`, where `𝕜` is a normed commutative ring and diff --git a/src/category_theory/Fintype.lean b/src/category_theory/Fintype.lean index 49fd297aeb310..08d43243895d2 100644 --- a/src/category_theory/Fintype.lean +++ b/src/category_theory/Fintype.lean @@ -13,6 +13,9 @@ import data.fintype.card /-! # The category of finite types. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define the category of finite types, denoted `Fintype` as (bundled) types with a `fintype` instance. diff --git a/src/category_theory/abelian/functor_category.lean b/src/category_theory/abelian/functor_category.lean index c8804931f2963..572d3f420b557 100644 --- a/src/category_theory/abelian/functor_category.lean +++ b/src/category_theory/abelian/functor_category.lean @@ -11,6 +11,9 @@ import category_theory.limits.preserves.shapes.kernels /-! # If `D` is abelian, then the functor category `C ⥤ D` is also abelian. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + -/ noncomputable theory diff --git a/src/category_theory/sites/plus.lean b/src/category_theory/sites/plus.lean index f286b413aecf4..b9673ffeb6709 100644 --- a/src/category_theory/sites/plus.lean +++ b/src/category_theory/sites/plus.lean @@ -9,6 +9,9 @@ import category_theory.sites.sheaf # The plus construction for presheaves. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains the construction of `P⁺`, for a presheaf `P : Cᵒᵖ ⥤ D` where `C` is endowed with a grothendieck topology `J`. diff --git a/src/category_theory/triangulated/pretriangulated.lean b/src/category_theory/triangulated/pretriangulated.lean index 4f101a20de6a3..bf6c5f20d6242 100644 --- a/src/category_theory/triangulated/pretriangulated.lean +++ b/src/category_theory/triangulated/pretriangulated.lean @@ -10,6 +10,9 @@ import category_theory.triangulated.rotate /-! # Pretriangulated Categories +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains the definition of pretriangulated categories and triangulated functors between them. diff --git a/src/category_theory/triangulated/rotate.lean b/src/category_theory/triangulated/rotate.lean index b4e212b14f6b9..748094c376071 100644 --- a/src/category_theory/triangulated/rotate.lean +++ b/src/category_theory/triangulated/rotate.lean @@ -9,6 +9,9 @@ import category_theory.triangulated.basic /-! # Rotate +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file adds the ability to rotate triangles and triangle morphisms. It also shows that rotation gives an equivalence on the category of triangles. diff --git a/src/category_theory/triangulated/triangulated.lean b/src/category_theory/triangulated/triangulated.lean index 5c78f2f72d905..67892425d0ee2 100644 --- a/src/category_theory/triangulated/triangulated.lean +++ b/src/category_theory/triangulated/triangulated.lean @@ -9,6 +9,9 @@ import category_theory.triangulated.pretriangulated /-! # Triangulated Categories +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains the definition of triangulated categories, which are pretriangulated categories which satisfy the octahedron axiom. diff --git a/src/control/fold.lean b/src/control/fold.lean index ad8ee814af62e..487b7193192d2 100644 --- a/src/control/fold.lean +++ b/src/control/fold.lean @@ -14,6 +14,9 @@ import category_theory.category.Kleisli # List folds generalized to `traversable` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Informally, we can think of `foldl` as a special case of `traverse` where we do not care about the reconstructed data structure and, in a state monad, we care about the final state. diff --git a/src/data/finset/sups.lean b/src/data/finset/sups.lean index 6928f2d00bc6e..0c2c69a056427 100644 --- a/src/data/finset/sups.lean +++ b/src/data/finset/sups.lean @@ -9,6 +9,9 @@ import data.set.sups /-! # Set family operations +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines a few binary operations on `finset α` for use in set family combinatorics. ## Main declarations diff --git a/src/data/list/to_finsupp.lean b/src/data/list/to_finsupp.lean index 3de6de022ed0f..be6a39ccf4be5 100644 --- a/src/data/list/to_finsupp.lean +++ b/src/data/list/to_finsupp.lean @@ -10,6 +10,9 @@ import data.finsupp.basic # Lists as finsupp +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + # Main definitions - `list.to_finsupp`: Interpret a list as a finitely supported function, where the indexing type diff --git a/src/data/num/lemmas.lean b/src/data/num/lemmas.lean index bfbfe1449518f..0829d392077b4 100644 --- a/src/data/num/lemmas.lean +++ b/src/data/num/lemmas.lean @@ -11,6 +11,9 @@ import data.nat.size /-! # Properties of the binary representation of integers + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ local attribute [simp] add_assoc diff --git a/src/field_theory/finiteness.lean b/src/field_theory/finiteness.lean index 39f8634bdbc00..55746cd353899 100644 --- a/src/field_theory/finiteness.lean +++ b/src/field_theory/finiteness.lean @@ -9,6 +9,9 @@ import linear_algebra.dimension /-! # A module over a division ring is noetherian if and only if it is finite. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + -/ universes u v diff --git a/src/field_theory/mv_polynomial.lean b/src/field_theory/mv_polynomial.lean index e048ed6a48c87..93d2ee62db8d8 100644 --- a/src/field_theory/mv_polynomial.lean +++ b/src/field_theory/mv_polynomial.lean @@ -12,6 +12,9 @@ import ring_theory.mv_polynomial.basic /-! # Multivariate polynomials over fields +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains basic facts about multivariate polynomials over fields, for example that the dimension of the space of multivariate polynomials over a field is equal to the cardinality of finitely supported functions from the indexing set to `ℕ`. diff --git a/src/linear_algebra/alternating.lean b/src/linear_algebra/alternating.lean index fdaee20a618eb..56101ce060757 100644 --- a/src/linear_algebra/alternating.lean +++ b/src/linear_algebra/alternating.lean @@ -14,6 +14,9 @@ import linear_algebra.multilinear.tensor_product /-! # Alternating Maps +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We construct the bundled function `alternating_map`, which extends `multilinear_map` with all the arguments of the same type. diff --git a/src/linear_algebra/dimension.lean b/src/linear_algebra/dimension.lean index ee150002b15e5..14c6dfd6c250d 100644 --- a/src/linear_algebra/dimension.lean +++ b/src/linear_algebra/dimension.lean @@ -14,6 +14,9 @@ import set_theory.cardinal.cofinality /-! # Dimension of modules and vector spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions * The rank of a module is defined as `module.rank : cardinal`. diff --git a/src/linear_algebra/free_module/finite/basic.lean b/src/linear_algebra/free_module/finite/basic.lean index c00cd6a61b431..e46f75d83b35e 100644 --- a/src/linear_algebra/free_module/finite/basic.lean +++ b/src/linear_algebra/free_module/finite/basic.lean @@ -11,6 +11,9 @@ import ring_theory.finiteness /-! # Finite and free modules +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We provide some instances for finite and free modules. ## Main results diff --git a/src/order/height.lean b/src/order/height.lean index 553a9744fc844..5e0a398afedd6 100644 --- a/src/order/height.lean +++ b/src/order/height.lean @@ -11,6 +11,9 @@ import tactic.tfae # Maximal length of chains +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains lemmas to work with the maximal length of strictly descending finite sequences (chains) in a partial order. diff --git a/src/set_theory/zfc/ordinal.lean b/src/set_theory/zfc/ordinal.lean index 0ff3f4855fc73..0a3080e266bb9 100644 --- a/src/set_theory/zfc/ordinal.lean +++ b/src/set_theory/zfc/ordinal.lean @@ -9,6 +9,9 @@ import set_theory.zfc.basic /-! # Von Neumann ordinals +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file works towards the development of von Neumann ordinals, i.e. transitive sets, well-ordered under `∈`. We currently only have an initial development of transitive sets. diff --git a/src/topology/algebra/module/weak_dual.lean b/src/topology/algebra/module/weak_dual.lean index 878a1b9f0b47a..0b4822cf8b21c 100644 --- a/src/topology/algebra/module/weak_dual.lean +++ b/src/topology/algebra/module/weak_dual.lean @@ -9,6 +9,9 @@ import linear_algebra.bilinear_map /-! # Weak dual topology +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the weak topology given two vector spaces `E` and `F` over a commutative semiring `𝕜` and a bilinear form `B : E →ₗ[𝕜] F →ₗ[𝕜] 𝕜`. The weak topology on `E` is the coarsest topology such that for all `y : F` every map `λ x, B x y` is continuous. diff --git a/src/topology/local_at_target.lean b/src/topology/local_at_target.lean index 9a32dd7651a53..9b8b6f5595f81 100644 --- a/src/topology/local_at_target.lean +++ b/src/topology/local_at_target.lean @@ -8,6 +8,9 @@ import topology.sets.opens /-! # Properties of maps that are local at the target. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We show that the following properties of continuous maps are local at the target : - `inducing` - `embedding` diff --git a/src/topology/metric_space/contracting.lean b/src/topology/metric_space/contracting.lean index 743f5fe0e4d63..daaac22d04149 100644 --- a/src/topology/metric_space/contracting.lean +++ b/src/topology/metric_space/contracting.lean @@ -10,6 +10,9 @@ import dynamics.fixed_points.topology /-! # Contracting maps +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A Lipschitz continuous self-map with Lipschitz constant `K < 1` is called a *contracting map*. In this file we prove the Banach fixed point theorem, some explicit estimates on the rate of convergence, and some properties of the map sending a contracting map to its fixed point. diff --git a/src/topology/metric_space/pi_nat.lean b/src/topology/metric_space/pi_nat.lean index d1ff39260c202..b3e2176c72d63 100644 --- a/src/topology/metric_space/pi_nat.lean +++ b/src/topology/metric_space/pi_nat.lean @@ -9,6 +9,9 @@ import topology.metric_space.hausdorff_distance /-! # Topological study of spaces `Π (n : ℕ), E n` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + When `E n` are topological spaces, the space `Π (n : ℕ), E n` is naturally a topological space (with the product topology). When `E n` are uniform spaces, it also inherits a uniform structure. However, it does not inherit a canonical metric space structure of the `E n`. Nevertheless, one From 23b80727b34e571e2e3bd8e8b720820cb215e880 Mon Sep 17 00:00:00 2001 From: themathqueen <2497250a@research.gla.ac.uk> Date: Tue, 11 Apr 2023 12:03:05 +0000 Subject: [PATCH 0810/1291] chore(analysis/inner_product_space/symmetric): change lemma name (#18777) changed name from [`linear_map.is_symmetric.inner_map_eq_zero`](https://leanprover-community.github.io/mathlib_docs/analysis/inner_product_space/symmetric.html#linear_map.is_symmetric.inner_map_eq_zero) to `linear_map.is_symmetric.inner_map_self_eq_zero` to match [`inner_map_self_eq_zero`](https://leanprover-community.github.io/mathlib_docs/analysis/inner_product_space/basic.html#inner_map_self_eq_zero) --- src/analysis/inner_product_space/symmetric.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/analysis/inner_product_space/symmetric.lean b/src/analysis/inner_product_space/symmetric.lean index c86dd6988052b..51143dd81b7b9 100644 --- a/src/analysis/inner_product_space/symmetric.lean +++ b/src/analysis/inner_product_space/symmetric.lean @@ -178,7 +178,7 @@ end /-- A symmetric linear map `T` is zero if and only if `⟪T x, x⟫_ℝ = 0` for all `x`. See `inner_map_self_eq_zero` for the complex version without the symmetric assumption. -/ -lemma is_symmetric.inner_map_eq_zero {T : E →ₗ[𝕜] E} (hT : T.is_symmetric) : +lemma is_symmetric.inner_map_self_eq_zero {T : E →ₗ[𝕜] E} (hT : T.is_symmetric) : (∀ x, ⟪T x, x⟫ = 0) ↔ T = 0 := begin simp_rw [linear_map.ext_iff, zero_apply], From 039ef89bef6e58b32b62898dd48e9d1a4312bb65 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 11 Apr 2023 13:21:35 +0000 Subject: [PATCH 0811/1291] chore(ring_theory/finiteness): generalize `finite_dimensional_range` to modules (#18787) --- src/linear_algebra/finite_dimensional.lean | 2 +- src/ring_theory/finiteness.lean | 4 ++++ 2 files changed, 5 insertions(+), 1 deletion(-) diff --git a/src/linear_algebra/finite_dimensional.lean b/src/linear_algebra/finite_dimensional.lean index 285285d308288..b72a40b849e35 100644 --- a/src/linear_algebra/finite_dimensional.lean +++ b/src/linear_algebra/finite_dimensional.lean @@ -875,7 +875,7 @@ 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 [finite_dimensional K V] (f : V →ₗ[K] V₂) : finite_dimensional K f.range := -f.quot_ker_equiv_range.finite_dimensional +module.finite.range f /-- 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} : diff --git a/src/ring_theory/finiteness.lean b/src/ring_theory/finiteness.lean index bc77fbe165d21..4f3d102acda03 100644 --- a/src/ring_theory/finiteness.lean +++ b/src/ring_theory/finiteness.lean @@ -475,6 +475,10 @@ lemma of_surjective [hM : finite R M] (f : M →ₗ[R] N) (hf : surjective f) : exact hM.1.map f end⟩ +/-- The range of a linear map from a finite module is finite. -/ +instance range [finite R M] (f : M →ₗ[R] N) : finite R f.range := +of_surjective f.range_restrict $ λ ⟨x, y, hy⟩, ⟨y, subtype.ext hy⟩ + variables (R) instance self : finite R R := From 46822d96c0c0a8e58a41a0cba1291620967575c5 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 12 Apr 2023 06:14:59 +0000 Subject: [PATCH 0812/1291] feat(linear_algebra/matrix/dot_product): `dot_product_self_eq_zero` (#18783) --- src/linear_algebra/matrix/dot_product.lean | 30 +++++++++++++++++++++- 1 file changed, 29 insertions(+), 1 deletion(-) diff --git a/src/linear_algebra/matrix/dot_product.lean b/src/linear_algebra/matrix/dot_product.lean index 28b475fdb742c..02f5afd931578 100644 --- a/src/linear_algebra/matrix/dot_product.lean +++ b/src/linear_algebra/matrix/dot_product.lean @@ -30,10 +30,12 @@ matrix, reindex -/ universes v w +variables {R : Type v} {n : Type w} namespace matrix -variables {R : Type v} [semiring R] {n : Type w} [fintype n] +section semiring +variables [semiring R] [fintype n] @[simp] lemma dot_product_std_basis_eq_mul [decidable_eq n] (v : n → R) (c : R) (i : n) : dot_product v (linear_map.std_basis R (λ _, R) i c) = v i * c := @@ -65,4 +67,30 @@ dot_product_eq _ _ $ λ u, (h u).symm ▸ (zero_dot_product u).symm lemma dot_product_eq_zero_iff {v : n → R} : (∀ w, dot_product v w = 0) ↔ v = 0 := ⟨λ h, dot_product_eq_zero v h, λ h w, h.symm ▸ zero_dot_product w⟩ +end semiring + +section self +variables [fintype n] + +@[simp] lemma dot_product_self_eq_zero [linear_ordered_ring R] {v : n → R} : + dot_product v v = 0 ↔ v = 0 := +(finset.sum_eq_zero_iff_of_nonneg $ λ i _, mul_self_nonneg (v i)).trans $ + by simp [function.funext_iff] + +/-- Note that this applies to `ℂ` via `complex.strict_ordered_comm_ring`. -/ +@[simp] lemma dot_product_star_self_eq_zero + [strict_ordered_ring R] [star_ordered_ring R] [no_zero_divisors R] {v : n → R} : + dot_product (star v) v = 0 ↔ v = 0 := +(finset.sum_eq_zero_iff_of_nonneg $ λ i _, @star_mul_self_nonneg _ _ _ _ (v i)).trans $ + by simp [function.funext_iff, mul_eq_zero] + +/-- Note that this applies to `ℂ` via `complex.strict_ordered_comm_ring`. -/ +@[simp] lemma dot_product_self_star_eq_zero + [strict_ordered_ring R] [star_ordered_ring R] [no_zero_divisors R] {v : n → R} : + dot_product v (star v) = 0 ↔ v = 0 := +(finset.sum_eq_zero_iff_of_nonneg $ λ i _, @star_mul_self_nonneg' _ _ _ _ (v i)).trans $ + by simp [function.funext_iff, mul_eq_zero] + +end self + end matrix From 9bb28972724354ac0574e2b318be896ec252025f Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Wed, 12 Apr 2023 07:22:21 +0000 Subject: [PATCH 0813/1291] chore(set_theory/cardinal/basic): reinstate `partial_order` (#18781) This change from #18714 was causing computability issues in Lean 4. Co-authored-by: Eric Wieser --- src/set_theory/cardinal/basic.lean | 21 ++++++++++++++++----- 1 file changed, 16 insertions(+), 5 deletions(-) diff --git a/src/set_theory/cardinal/basic.lean b/src/set_theory/cardinal/basic.lean index f7eb3ef0a70a5..d3d08052171f7 100644 --- a/src/set_theory/cardinal/basic.lean +++ b/src/set_theory/cardinal/basic.lean @@ -191,13 +191,16 @@ instance : has_le cardinal.{u} := ⟨λ q₁ q₂, quotient.lift_on₂ q₁ q₂ (λ α β, nonempty $ α ↪ β) $ λ α β γ δ ⟨e₁⟩ ⟨e₂⟩, propext ⟨λ ⟨e⟩, ⟨e.congr e₁ e₂⟩, λ ⟨e⟩, ⟨e.congr e₁.symm e₂.symm⟩⟩⟩ -instance : linear_order cardinal.{u} := +instance : partial_order cardinal.{u} := { le := (≤), le_refl := by rintros ⟨α⟩; exact ⟨embedding.refl _⟩, le_trans := by rintros ⟨α⟩ ⟨β⟩ ⟨γ⟩ ⟨e₁⟩ ⟨e₂⟩; exact ⟨e₁.trans e₂⟩, - le_antisymm := by { rintros ⟨α⟩ ⟨β⟩ ⟨e₁⟩ ⟨e₂⟩, exact quotient.sound (e₁.antisymm e₂) }, - le_total := by { rintros ⟨α⟩ ⟨β⟩, apply embedding.total }, - decidable_le := classical.dec_rel _ } + le_antisymm := by { rintros ⟨α⟩ ⟨β⟩ ⟨e₁⟩ ⟨e₂⟩, exact quotient.sound (e₁.antisymm e₂) }, } + +instance : linear_order cardinal.{u} := +{ le_total := by { rintros ⟨α⟩ ⟨β⟩, apply embedding.total }, + decidable_le := classical.dec_rel _, + ..cardinal.partial_order.{u}, } theorem le_def (α β : Type u) : #α ≤ #β ↔ nonempty (α ↪ β) := iff.rfl @@ -475,18 +478,26 @@ instance : canonically_ordered_comm_semiring cardinal.{u} := le_self_add := λ a b, (add_zero a).ge.trans $ add_le_add_left (cardinal.zero_le _) _, eq_zero_or_eq_zero_of_mul_eq_zero := λ a b, induction_on₂ a b $ λ α β, by simpa only [mul_def, mk_eq_zero_iff, is_empty_prod] using id, - ..cardinal.comm_semiring, ..cardinal.linear_order } + ..cardinal.comm_semiring, ..cardinal.partial_order } instance : canonically_linear_ordered_add_monoid cardinal.{u} := { ..cardinal.canonically_ordered_comm_semiring, ..cardinal.linear_order } +-- Computable instance to prevent a non-computable one being found via the one above +instance : canonically_ordered_add_monoid cardinal.{u} := +{ ..cardinal.canonically_ordered_comm_semiring } + instance : linear_ordered_comm_monoid_with_zero cardinal.{u} := { mul_le_mul_left := @mul_le_mul_left' _ _ _ _, zero_le_one := zero_le _, ..cardinal.comm_semiring, ..cardinal.linear_order } +-- Computable instance to prevent a non-computable one being found via the one above +instance : comm_monoid_with_zero cardinal.{u} := +{ ..cardinal.canonically_ordered_comm_semiring } + lemma zero_power_le (c : cardinal.{u}) : (0 : cardinal.{u}) ^ c ≤ 1 := by { by_cases h : c = 0, rw [h, power_zero], rw [zero_power h], apply zero_le } From 814d76e2247d5ba8bc024843552da1278bfe9e5c Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Wed, 12 Apr 2023 07:22:22 +0000 Subject: [PATCH 0814/1291] chore(*): add mathlib4 synchronization comments (#18793) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.category.Module.basic` * `algebra.linear_recurrence` * `algebra.order.chebyshev` * `algebra.order.to_interval_mod` * `algebraic_topology.simplicial_object` * `analysis.convex.extrema` * `analysis.convex.slope` * `analysis.normed.order.basic` * `analysis.normed_space.basic` * `category_theory.graded_object` * `computability.encoding` * `geometry.manifold.charted_space` * `linear_algebra.free_algebra` * `linear_algebra.quotient_pi` * `ring_theory.polynomial.eisenstein.basic` * `topology.category.Top.adjunctions` * `topology.category.Top.basic` --- src/algebra/category/Module/basic.lean | 3 +++ src/algebra/linear_recurrence.lean | 3 +++ src/algebra/order/chebyshev.lean | 3 +++ src/algebra/order/to_interval_mod.lean | 3 +++ src/algebraic_topology/simplicial_object.lean | 3 +++ src/analysis/convex/extrema.lean | 3 +++ src/analysis/convex/slope.lean | 3 +++ src/analysis/normed/order/basic.lean | 3 +++ src/analysis/normed_space/basic.lean | 3 +++ src/category_theory/graded_object.lean | 3 +++ src/computability/encoding.lean | 3 +++ src/geometry/manifold/charted_space.lean | 3 +++ src/linear_algebra/free_algebra.lean | 3 +++ src/linear_algebra/quotient_pi.lean | 3 +++ src/ring_theory/polynomial/eisenstein/basic.lean | 3 +++ src/topology/category/Top/adjunctions.lean | 3 +++ src/topology/category/Top/basic.lean | 3 +++ 17 files changed, 51 insertions(+) diff --git a/src/algebra/category/Module/basic.lean b/src/algebra/category/Module/basic.lean index ddd4a0f5691c1..72ea05ddfb8ad 100644 --- a/src/algebra/category/Module/basic.lean +++ b/src/algebra/category/Module/basic.lean @@ -13,6 +13,9 @@ import category_theory.preadditive.additive_functor /-! # The category of `R`-modules +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + `Module.{v} R` is the category of bundled `R`-modules with carrier in the universe `v`. We show that it is preadditive and show that being an isomorphism, monomorphism and epimorphism is equivalent to being a linear equivalence, an injective linear map and a surjective linear map, diff --git a/src/algebra/linear_recurrence.lean b/src/algebra/linear_recurrence.lean index e1e570166c3ae..1e9b19c88f546 100644 --- a/src/algebra/linear_recurrence.lean +++ b/src/algebra/linear_recurrence.lean @@ -9,6 +9,9 @@ import linear_algebra.dimension /-! # Linear recurrence +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Informally, a "linear recurrence" is an assertion of the form `∀ n : ℕ, u (n + d) = a 0 * u n + a 1 * u (n+1) + ... + a (d-1) * u (n+d-1)`, where `u` is a sequence, `d` is the *order* of the recurrence and the `a i` diff --git a/src/algebra/order/chebyshev.lean b/src/algebra/order/chebyshev.lean index 6a1bdef1eae39..619b8b6c0d5ed 100644 --- a/src/algebra/order/chebyshev.lean +++ b/src/algebra/order/chebyshev.lean @@ -10,6 +10,9 @@ import group_theory.perm.cycle.basic /-! # Chebyshev's sum inequality +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves the Chebyshev sum inequality. Chebyshev's inequality states `(∑ i in s, f i) * (∑ i in s, g i) ≤ s.card * ∑ i in s, f i * g i` diff --git a/src/algebra/order/to_interval_mod.lean b/src/algebra/order/to_interval_mod.lean index b153c053b0aa4..34a4176ebf5b6 100644 --- a/src/algebra/order/to_interval_mod.lean +++ b/src/algebra/order/to_interval_mod.lean @@ -12,6 +12,9 @@ import group_theory.quotient_group /-! # Reducing to an interval modulo its length +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines operations that reduce a number (in an `archimedean` `linear_ordered_add_comm_group`) to a number in a given interval, modulo the length of that interval. diff --git a/src/algebraic_topology/simplicial_object.lean b/src/algebraic_topology/simplicial_object.lean index 71c21b0de39fe..0467fe961b36d 100644 --- a/src/algebraic_topology/simplicial_object.lean +++ b/src/algebraic_topology/simplicial_object.lean @@ -11,6 +11,9 @@ import category_theory.opposites /-! # Simplicial objects in a category. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A simplicial object in a category `C` is a `C`-valued presheaf on `simplex_category`. (Similarly a cosimplicial object is functor `simplex_category ⥤ C`.) diff --git a/src/analysis/convex/extrema.lean b/src/analysis/convex/extrema.lean index c5ee65948f40c..416e8a5a1fdcd 100644 --- a/src/analysis/convex/extrema.lean +++ b/src/analysis/convex/extrema.lean @@ -11,6 +11,9 @@ import topology.metric_space.basic /-! # Minima and maxima of convex functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We show that if a function `f : E → β` is convex, then a local minimum is also a global minimum, and likewise for concave functions. -/ diff --git a/src/analysis/convex/slope.lean b/src/analysis/convex/slope.lean index b60a692638e25..e6ddba269e7dc 100644 --- a/src/analysis/convex/slope.lean +++ b/src/analysis/convex/slope.lean @@ -10,6 +10,9 @@ import tactic.linarith /-! # Slopes of convex functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file relates convexity/concavity of functions in a linearly ordered field and the monotonicity of their slopes. diff --git a/src/analysis/normed/order/basic.lean b/src/analysis/normed/order/basic.lean index 92b1bb9f753cd..79e3daec9d62a 100644 --- a/src/analysis/normed/order/basic.lean +++ b/src/analysis/normed/order/basic.lean @@ -9,6 +9,9 @@ import analysis.normed_space.basic /-! # Ordered normed spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we define classes for fields and groups that are both normed and ordered. These are mostly useful to avoid diamonds during type class inference. -/ diff --git a/src/analysis/normed_space/basic.lean b/src/analysis/normed_space/basic.lean index ee8c4968a9851..1b623e6761747 100644 --- a/src/analysis/normed_space/basic.lean +++ b/src/analysis/normed_space/basic.lean @@ -12,6 +12,9 @@ import topology.algebra.module.basic /-! # Normed spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define (semi)normed spaces and algebras. We also prove some theorems about these definitions. -/ diff --git a/src/category_theory/graded_object.lean b/src/category_theory/graded_object.lean index 2ad314d225879..5bc8805c812c1 100644 --- a/src/category_theory/graded_object.lean +++ b/src/category_theory/graded_object.lean @@ -11,6 +11,9 @@ import category_theory.concrete_category.basic /-! # The category of graded objects +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + For any type `β`, a `β`-graded object over some category `C` is just a function `β → C` into the objects of `C`. We put the "pointwise" category structure on these, as the non-dependent specialization of diff --git a/src/computability/encoding.lean b/src/computability/encoding.lean index 77a1618662df1..f161ff7a6237c 100644 --- a/src/computability/encoding.lean +++ b/src/computability/encoding.lean @@ -12,6 +12,9 @@ import tactic.derive_fintype /-! # Encodings +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains the definition of a (finite) encoding, a map from a type to strings in an alphabet, used in defining computability by Turing machines. It also contains several examples: diff --git a/src/geometry/manifold/charted_space.lean b/src/geometry/manifold/charted_space.lean index 4f25641a22c03..519ab352e5316 100644 --- a/src/geometry/manifold/charted_space.lean +++ b/src/geometry/manifold/charted_space.lean @@ -8,6 +8,9 @@ import topology.local_homeomorph /-! # Charted spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A smooth manifold is a topological space `M` locally modelled on a euclidean space (or a euclidean half-space for manifolds with boundaries, or an infinite dimensional vector space for more general notions of manifolds), i.e., the manifold is covered by open subsets on which there are local diff --git a/src/linear_algebra/free_algebra.lean b/src/linear_algebra/free_algebra.lean index 409c4c0c8154a..846d10b1f393c 100644 --- a/src/linear_algebra/free_algebra.lean +++ b/src/linear_algebra/free_algebra.lean @@ -10,6 +10,9 @@ import linear_algebra.finsupp_vector_space /-! # Linear algebra properties of `free_algebra R X` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file provides a `free_monoid X` basis on the `free_algebra R X`, and uses it to show the dimension of the algebra is the cardinality of `list X` -/ diff --git a/src/linear_algebra/quotient_pi.lean b/src/linear_algebra/quotient_pi.lean index da66bef617d8b..7ec75b10ca9a4 100644 --- a/src/linear_algebra/quotient_pi.lean +++ b/src/linear_algebra/quotient_pi.lean @@ -9,6 +9,9 @@ import linear_algebra.quotient /-! # Submodule quotients and direct sums +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains some results on the quotient of a module by a direct sum of submodules, and the direct sum of quotients of modules by submodules. diff --git a/src/ring_theory/polynomial/eisenstein/basic.lean b/src/ring_theory/polynomial/eisenstein/basic.lean index 6d37190e8da92..3e3903a29d390 100644 --- a/src/ring_theory/polynomial/eisenstein/basic.lean +++ b/src/ring_theory/polynomial/eisenstein/basic.lean @@ -9,6 +9,9 @@ import ring_theory.polynomial.scale_roots /-! # Eisenstein polynomials + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. Given an ideal `𝓟` of a commutative semiring `R`, we say that a polynomial `f : R[X]` is *Eisenstein at `𝓟`* if `f.leading_coeff ∉ 𝓟`, `∀ n, n < f.nat_degree → f.coeff n ∈ 𝓟` and `f.coeff 0 ∉ 𝓟 ^ 2`. In this file we gather miscellaneous results about Eisenstein polynomials. diff --git a/src/topology/category/Top/adjunctions.lean b/src/topology/category/Top/adjunctions.lean index a0ce2a4cf7430..5083d17585d7c 100644 --- a/src/topology/category/Top/adjunctions.lean +++ b/src/topology/category/Top/adjunctions.lean @@ -9,6 +9,9 @@ import category_theory.adjunction.basic /-! # Adjunctions regarding the category of topological spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file shows that the forgetful functor from topological spaces to types has a left and right adjoint, given by `Top.discrete`, resp. `Top.trivial`, the functors which equip a type with the discrete, resp. trivial, topology. diff --git a/src/topology/category/Top/basic.lean b/src/topology/category/Top/basic.lean index 5a2da0ff45f5c..f212483491efe 100644 --- a/src/topology/category/Top/basic.lean +++ b/src/topology/category/Top/basic.lean @@ -10,6 +10,9 @@ import topology.continuous_function.basic /-! # Category instance for topological spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We introduce the bundled category `Top` of topological spaces together with the functors `discrete` and `trivial` from the category of types to `Top` which equip a type with the corresponding discrete, resp. trivial, topology. For a proof that these functors are left, resp. right adjoint From b5b5dd5a47b5744260e2c9185013075ce9dadccd Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 12 Apr 2023 09:52:25 +0000 Subject: [PATCH 0815/1291] feat(linear_algebra/free_module/finite/rank): remove `module.free` assumption (#18792) Combined with the result in #18787, this lets us golf a downstream proof about matrices. --- src/data/matrix/rank.lean | 7 ++-- src/linear_algebra/dimension.lean | 6 ++++ .../free_module/finite/rank.lean | 35 ++++++++++++------- 3 files changed, 30 insertions(+), 18 deletions(-) diff --git a/src/data/matrix/rank.lean b/src/data/matrix/rank.lean index f112526ca70a7..980ef1437ddd6 100644 --- a/src/data/matrix/rank.lean +++ b/src/data/matrix/rank.lean @@ -56,11 +56,8 @@ lemma rank_mul_le [strong_rank_condition R] (A : matrix m n R) (B : matrix n o R (A ⬝ B).rank ≤ A.rank := begin rw [rank, rank, to_lin'_mul], - refine cardinal.to_nat_le_of_le_of_lt_aleph_0 _ (linear_map.rank_comp_le_left _ _), - rw [←cardinal.lift_lt_aleph_0], - have := lift_rank_range_le A.to_lin', - rw [rank_fun', cardinal.lift_nat_cast] at this, - exact this.trans_lt (cardinal.nat_lt_aleph_0 (fintype.card n)), + exact cardinal.to_nat_le_of_le_of_lt_aleph_0 + (rank_lt_aleph_0 _ _) (linear_map.rank_comp_le_left _ _), end lemma rank_unit [strong_rank_condition R] (A : (matrix n n R)ˣ) : diff --git a/src/linear_algebra/dimension.lean b/src/linear_algebra/dimension.lean index 14c6dfd6c250d..4ba29e9dd770b 100644 --- a/src/linear_algebra/dimension.lean +++ b/src/linear_algebra/dimension.lean @@ -733,6 +733,12 @@ begin exact le_top, end +/-- A version of `linear_independent_le_span` for `finset`. -/ +lemma linear_independent_le_span_finset {ι : Type*} (v : ι → M) (i : linear_independent R v) + (w : finset M) (s : span R (w : set M) = ⊤) : + #ι ≤ w.card := +by simpa only [finset.coe_sort_coe, fintype.card_coe] using linear_independent_le_span v i w s + /-- An auxiliary lemma for `linear_independent_le_basis`: we handle the case where the basis `b` is infinite. diff --git a/src/linear_algebra/free_module/finite/rank.lean b/src/linear_algebra/free_module/finite/rank.lean index b4160c3f448e2..bc0b31beaddc5 100644 --- a/src/linear_algebra/free_module/finite/rank.lean +++ b/src/linear_algebra/free_module/finite/rank.lean @@ -32,21 +32,31 @@ open module.free section ring variables [ring R] [strong_rank_condition R] -variables [add_comm_group M] [module R M] [module.free R M] [module.finite R M] -variables [add_comm_group N] [module R N] [module.free R N] [module.finite R N] +variables [add_comm_group M] [module R M] [module.finite R M] +variables [add_comm_group N] [module R N] [module.finite R N] -/-- The rank of a finite and free module is finite. -/ +/-- The rank of a finite module is finite. -/ lemma rank_lt_aleph_0 : module.rank R M < ℵ₀ := begin + dunfold module.rank, letI := nontrivial_of_invariant_basis_number R, - rw [← (choose_basis R M).mk_eq_rank'', lt_aleph_0_iff_fintype], - exact nonempty.intro infer_instance + obtain ⟨S, hS⟩ := module.finite_def.mp ‹_›, + refine (csupr_le' $ λ i, _).trans_lt (nat_lt_aleph_0 S.card), + exact linear_independent_le_span_finset _ i.prop S hS, end -/-- If `M` is finite and free, `finrank M = rank M`. -/ +/-- If `M` is finite, `finrank M = rank M`. -/ @[simp] lemma finrank_eq_rank : ↑(finrank R M) = module.rank R M := by { rw [finrank, cast_to_nat_of_lt_aleph_0 (rank_lt_aleph_0 R M)] } +end ring + +section ring_free + +variables [ring R] [strong_rank_condition R] +variables [add_comm_group M] [module R M] [module.free R M] [module.finite R M] +variables [add_comm_group N] [module R N] [module.free R N] [module.finite R N] + /-- The finrank of a free module `M` over `R` is the cardinality of `choose_basis_index R M`. -/ lemma finrank_eq_card_choose_basis_index : finrank R M = @card (choose_basis_index R M) (@choose_basis_index.fintype R M _ _ _ _ (nontrivial_of_invariant_basis_number R) _) := @@ -94,7 +104,7 @@ lemma finrank_matrix (m n : Type*) [fintype m] [fintype n] : finrank R (matrix m n R) = (card m) * (card n) := by { simp [finrank] } -end ring +end ring_free section comm_ring @@ -120,24 +130,23 @@ variables [ring R] [strong_rank_condition R] variables [add_comm_group M] [module R M] variables [add_comm_group N] [module R N] -lemma linear_map.finrank_le_finrank_of_injective - [module.free R N] [module.finite R N] {f : M →ₗ[R] N} (hf : function.injective f) : - finrank R M ≤ finrank R N := +lemma linear_map.finrank_le_finrank_of_injective [module.finite R N] {f : M →ₗ[R] N} + (hf : function.injective f) : finrank R M ≤ finrank R N := finrank_le_finrank_of_rank_le_rank (linear_map.lift_rank_le_of_injective _ hf) (rank_lt_aleph_0 _ _) -lemma linear_map.finrank_range_le [module.free R M] [module.finite R M] (f : M →ₗ[R] N) : +lemma linear_map.finrank_range_le [module.finite R M] (f : M →ₗ[R] N) : finrank R f.range ≤ finrank R M := finrank_le_finrank_of_rank_le_rank (lift_rank_range_le f) (rank_lt_aleph_0 _ _) /-- The dimension of a submodule is bounded by the dimension of the ambient space. -/ -lemma submodule.finrank_le [module.free R M] [module.finite R M] (s : submodule R M) : +lemma submodule.finrank_le [module.finite R M] (s : submodule R M) : finrank R s ≤ finrank R M := by simpa only [cardinal.to_nat_lift] using to_nat_le_of_le_of_lt_aleph_0 (rank_lt_aleph_0 _ _) (rank_submodule_le s) /-- The dimension of a quotient is bounded by the dimension of the ambient space. -/ -lemma submodule.finrank_quotient_le [module.free R M] [module.finite R M] (s : submodule R M) : +lemma submodule.finrank_quotient_le [module.finite R M] (s : submodule R M) : finrank R (M ⧸ s) ≤ finrank R M := by simpa only [cardinal.to_nat_lift] using to_nat_le_of_le_of_lt_aleph_0 (rank_lt_aleph_0 _ _) ((submodule.mkq s).rank_le_of_surjective (surjective_quot_mk _)) From 347636a7a80595d55bedf6e6fbd996a3c39da69a Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Wed, 12 Apr 2023 15:48:11 +0000 Subject: [PATCH 0816/1291] chore(linear_algebra/finrank): backport removal of simp lemmas (#18794) Testing a solution to the simpNF linter problems at https://github.com/leanprover-community/mathlib4/pull/3378 Co-authored-by: Scott Morrison --- src/linear_algebra/finrank.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/linear_algebra/finrank.lean b/src/linear_algebra/finrank.lean index b78c2a2f8dfdd..70917e4dfafbb 100644 --- a/src/linear_algebra/finrank.lean +++ b/src/linear_algebra/finrank.lean @@ -432,7 +432,7 @@ basis.mk (linear_independent_of_top_le_span_of_card_eq_finrank le_span card_eq) basis.coe_mk _ _ /-- A finset of `finrank K V` vectors forms a basis if they span the whole space. -/ -@[simps] +@[simps repr_apply] noncomputable def finset_basis_of_top_le_span_of_card_eq_finrank {s : finset V} (le_span : ⊤ ≤ span K (s : set V)) (card_eq : s.card = finrank K V) : basis (s : set V) K V := @@ -441,7 +441,7 @@ basis_of_top_le_span_of_card_eq_finrank (coe : (s : set V) → V) (trans (fintype.card_coe _) card_eq) /-- A set of `finrank K V` vectors forms a basis if they span the whole space. -/ -@[simps] +@[simps repr_apply] noncomputable def set_basis_of_top_le_span_of_card_eq_finrank {s : set V} [fintype s] (le_span : ⊤ ≤ span K s) (card_eq : s.to_finset.card = finrank K V) : basis s K V := From 5ac1dab1670014b4c07a82c86a67f3d064a1b3e1 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 12 Apr 2023 23:51:56 +0000 Subject: [PATCH 0817/1291] chore(linear_algebra/matrix/dot_product): weaken typeclasses (#18798) This makes unification slightly harder on Lean, so `: _`s are added. Fixes a stupid error in #18783 --- src/linear_algebra/matrix/dot_product.lean | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/src/linear_algebra/matrix/dot_product.lean b/src/linear_algebra/matrix/dot_product.lean index 02f5afd931578..1aca66ab84fd6 100644 --- a/src/linear_algebra/matrix/dot_product.lean +++ b/src/linear_algebra/matrix/dot_product.lean @@ -79,16 +79,16 @@ variables [fintype n] /-- Note that this applies to `ℂ` via `complex.strict_ordered_comm_ring`. -/ @[simp] lemma dot_product_star_self_eq_zero - [strict_ordered_ring R] [star_ordered_ring R] [no_zero_divisors R] {v : n → R} : + [partial_order R] [non_unital_ring R] [star_ordered_ring R] [no_zero_divisors R] {v : n → R} : dot_product (star v) v = 0 ↔ v = 0 := -(finset.sum_eq_zero_iff_of_nonneg $ λ i _, @star_mul_self_nonneg _ _ _ _ (v i)).trans $ +(finset.sum_eq_zero_iff_of_nonneg $ λ i _, (@star_mul_self_nonneg _ _ _ _ (v i) : _)).trans $ by simp [function.funext_iff, mul_eq_zero] /-- Note that this applies to `ℂ` via `complex.strict_ordered_comm_ring`. -/ @[simp] lemma dot_product_self_star_eq_zero - [strict_ordered_ring R] [star_ordered_ring R] [no_zero_divisors R] {v : n → R} : + [partial_order R] [non_unital_ring R] [star_ordered_ring R] [no_zero_divisors R] {v : n → R} : dot_product v (star v) = 0 ↔ v = 0 := -(finset.sum_eq_zero_iff_of_nonneg $ λ i _, @star_mul_self_nonneg' _ _ _ _ (v i)).trans $ +(finset.sum_eq_zero_iff_of_nonneg $ λ i _, (@star_mul_self_nonneg' _ _ _ _ (v i) : _)).trans $ by simp [function.funext_iff, mul_eq_zero] end self From 21407283bcca2ef7e3c7df50bd8bb42179f023b3 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 13 Apr 2023 05:16:58 +0000 Subject: [PATCH 0818/1291] feat(data/matrix/rank): rank of a matrix is the rank of its column space (#18778) This is a TODO comment in this file left from #10826. Proving the link to the row space is harder, and only easy to prove for `is_R_or_C`; so I've left it for another time. --- src/data/matrix/rank.lean | 7 ++++++- 1 file changed, 6 insertions(+), 1 deletion(-) diff --git a/src/data/matrix/rank.lean b/src/data/matrix/rank.lean index 980ef1437ddd6..bed691aec2208 100644 --- a/src/data/matrix/rank.lean +++ b/src/data/matrix/rank.lean @@ -19,7 +19,7 @@ This definition does not depend on the choice of basis, see `matrix.rank_eq_finr ## TODO -* Show that `matrix.rank` is equal to the row-rank and column-rank +* Show that `matrix.rank` is equal to the row-rank, and that `rank Aᵀ = rank A`. -/ @@ -110,4 +110,9 @@ lemma rank_le_height [strong_rank_condition R] {m n : ℕ} (A : matrix (fin m) ( A.rank ≤ m := A.rank_le_card_height.trans $ (fintype.card_fin m).le +/-- The rank of a matrix is the rank of the space spanned by its columns. -/ +lemma rank_eq_finrank_span_cols (A : matrix m n R) : + A.rank = finrank R (submodule.span R (set.range Aᵀ)) := +by rw [rank, matrix.range_to_lin'] + end matrix From 9a48a083b390d9b84a71efbdc4e8dfa26a687104 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Thu, 13 Apr 2023 06:19:53 +0000 Subject: [PATCH 0819/1291] chore(*): add mathlib4 synchronization comments (#18799) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `analysis.asymptotics.asymptotics` * `analysis.convex.quasiconvex` * `analysis.locally_convex.basic` * `analysis.normed.order.upper_lower` * `analysis.normed_space.continuous_linear_map` * `analysis.normed_space.ray` * `analysis.normed_space.riesz_lemma` * `category_theory.idempotents.simplicial_object` * `group_theory.specific_groups.alternating` * `linear_algebra.finrank` * `model_theory.language_map` --- src/analysis/asymptotics/asymptotics.lean | 3 +++ src/analysis/convex/quasiconvex.lean | 3 +++ src/analysis/locally_convex/basic.lean | 3 +++ src/analysis/normed/order/upper_lower.lean | 3 +++ src/analysis/normed_space/continuous_linear_map.lean | 3 +++ src/analysis/normed_space/ray.lean | 3 +++ src/analysis/normed_space/riesz_lemma.lean | 3 +++ src/category_theory/idempotents/simplicial_object.lean | 3 +++ src/group_theory/specific_groups/alternating.lean | 3 +++ src/linear_algebra/finrank.lean | 3 +++ src/model_theory/language_map.lean | 3 +++ 11 files changed, 33 insertions(+) diff --git a/src/analysis/asymptotics/asymptotics.lean b/src/analysis/asymptotics/asymptotics.lean index 21eff942ac3c3..77c2586c1dc1d 100644 --- a/src/analysis/asymptotics/asymptotics.lean +++ b/src/analysis/asymptotics/asymptotics.lean @@ -11,6 +11,9 @@ import topology.local_homeomorph /-! # Asymptotics +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We introduce these relations: * `is_O_with c l f g` : "f is big O of g along l with constant c"; diff --git a/src/analysis/convex/quasiconvex.lean b/src/analysis/convex/quasiconvex.lean index 351f7f7fd2222..5cdb6f4f8cd7e 100644 --- a/src/analysis/convex/quasiconvex.lean +++ b/src/analysis/convex/quasiconvex.lean @@ -8,6 +8,9 @@ import analysis.convex.function /-! # Quasiconvex and quasiconcave functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines quasiconvexity, quasiconcavity and quasilinearity of functions, which are generalizations of unimodality and monotonicity. Convexity implies quasiconvexity, concavity implies quasiconcavity, and monotonicity implies quasilinearity. diff --git a/src/analysis/locally_convex/basic.lean b/src/analysis/locally_convex/basic.lean index 39656c98883e0..49ce8cc61ba98 100644 --- a/src/analysis/locally_convex/basic.lean +++ b/src/analysis/locally_convex/basic.lean @@ -10,6 +10,9 @@ import analysis.normed_space.basic /-! # Local convexity +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines absorbent and balanced sets. An absorbent set is one that "surrounds" the origin. The idea is made precise by requiring that any diff --git a/src/analysis/normed/order/upper_lower.lean b/src/analysis/normed/order/upper_lower.lean index fa1fe8f25bf1a..c834d2d671ddc 100644 --- a/src/analysis/normed/order/upper_lower.lean +++ b/src/analysis/normed/order/upper_lower.lean @@ -11,6 +11,9 @@ import topology.algebra.order.upper_lower /-! # Upper/lower/order-connected sets in normed groups +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The topological closure and interior of an upper/lower/order-connected set is an upper/lower/order-connected set (with the notable exception of the closure of an order-connected set). diff --git a/src/analysis/normed_space/continuous_linear_map.lean b/src/analysis/normed_space/continuous_linear_map.lean index 50c0db7872b3f..693adffcd123e 100644 --- a/src/analysis/normed_space/continuous_linear_map.lean +++ b/src/analysis/normed_space/continuous_linear_map.lean @@ -7,6 +7,9 @@ import analysis.normed_space.basic /-! # Constructions of continuous linear maps between (semi-)normed spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A fundamental fact about (semi-)linear maps between normed spaces over sensible fields is that continuity and boundedness are equivalent conditions. That is, for normed spaces `E`, `F`, a `linear_map` `f : E →ₛₗ[σ] F` is the coercion of some `continuous_linear_map` `f' : E →SL[σ] F`, if diff --git a/src/analysis/normed_space/ray.lean b/src/analysis/normed_space/ray.lean index f2e63ce6ee08d..d6724816a61b4 100644 --- a/src/analysis/normed_space/ray.lean +++ b/src/analysis/normed_space/ray.lean @@ -9,6 +9,9 @@ import analysis.normed_space.basic /-! # Rays in a real normed vector space +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove some lemmas about the `same_ray` predicate in case of a real normed space. In this case, for two vectors `x y` in the same ray, the norm of their sum is equal to the sum of their norms and `‖y‖ • x = ‖x‖ • y`. diff --git a/src/analysis/normed_space/riesz_lemma.lean b/src/analysis/normed_space/riesz_lemma.lean index 4cc4cec092439..0e14a11353071 100644 --- a/src/analysis/normed_space/riesz_lemma.lean +++ b/src/analysis/normed_space/riesz_lemma.lean @@ -9,6 +9,9 @@ import topology.metric_space.hausdorff_distance /-! # Applications of the Hausdorff distance in normed spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Riesz's lemma, stated for a normed space over a normed field: for any closed proper subspace `F` of `E`, there is a nonzero `x` such that `‖x - F‖` is at least `r * ‖x‖` for any `r < 1`. This is `riesz_lemma`. diff --git a/src/category_theory/idempotents/simplicial_object.lean b/src/category_theory/idempotents/simplicial_object.lean index 0432a50a0e2ab..cd18779b827ee 100644 --- a/src/category_theory/idempotents/simplicial_object.lean +++ b/src/category_theory/idempotents/simplicial_object.lean @@ -11,6 +11,9 @@ import category_theory.idempotents.functor_categories # Idempotent completeness of categories of simplicial objects +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we provide an instance expressing that `simplicial_object C` and `cosimplicial_object C` are idempotent complete categories when the category `C` is. diff --git a/src/group_theory/specific_groups/alternating.lean b/src/group_theory/specific_groups/alternating.lean index 5f4a13d2e7245..89f6761f9cde4 100644 --- a/src/group_theory/specific_groups/alternating.lean +++ b/src/group_theory/specific_groups/alternating.lean @@ -12,6 +12,9 @@ import tactic.interval_cases /-! # Alternating Groups +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The alternating group on a finite type `α` is the subgroup of the permutation group `perm α` consisting of the even permutations. diff --git a/src/linear_algebra/finrank.lean b/src/linear_algebra/finrank.lean index 70917e4dfafbb..5ff06c9f2cf74 100644 --- a/src/linear_algebra/finrank.lean +++ b/src/linear_algebra/finrank.lean @@ -8,6 +8,9 @@ import linear_algebra.dimension /-! # Finite dimension of vector spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Definition of the rank of a module, or dimension of a vector space, as a natural number. ## Main definitions diff --git a/src/model_theory/language_map.lean b/src/model_theory/language_map.lean index 7407b211d7e5a..a6a5e85c3d534 100644 --- a/src/model_theory/language_map.lean +++ b/src/model_theory/language_map.lean @@ -6,6 +6,9 @@ Authors: Aaron Anderson, Jesse Michael Han, Floris van Doorn import model_theory.basic /-! # Language Maps + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. Maps between first-order languages in the style of the [Flypitch project](https://flypitch.github.io/), as well as several important maps between structures. From 47a5f8186becdbc826190ced4312f8199f9db6a5 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 13 Apr 2023 09:44:03 +0000 Subject: [PATCH 0820/1291] feat(data/matrix/rank): rank of multiplication (#18784) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This adds a universe-polymorphic version of `rank_comp_le_right`, and then uses it to show `(A ⬝ B).rank ≤ B.rank`; previously we only had `(A ⬝ B).rank ≤ A.rank`. For convenience, this adds the spellings `(A ⬝ B).rank ≤ min A.rank B.rank` and `rank (f.comp g) ≤ min (rank f) (rank g)`, as these map well to the way that rank would be described in words. --- src/data/matrix/rank.lean | 23 ++++++++++++++++++++--- src/linear_algebra/dimension.lean | 19 ++++++++++++++++++- 2 files changed, 38 insertions(+), 4 deletions(-) diff --git a/src/data/matrix/rank.lean b/src/data/matrix/rank.lean index bed691aec2208..db24000f23639 100644 --- a/src/data/matrix/rank.lean +++ b/src/data/matrix/rank.lean @@ -29,7 +29,7 @@ namespace matrix open finite_dimensional -variables {m n o R : Type*} [m_fin : fintype m] [fintype n] [fintype o] +variables {l m n o R : Type*} [m_fin : fintype m] [fintype n] [fintype o] variables [decidable_eq n] [decidable_eq o] [comm_ring R] /-- The rank of a matrix is the rank of its image. -/ @@ -52,7 +52,7 @@ lemma rank_le_width [strong_rank_condition R] {m n : ℕ} (A : matrix (fin m) (f A.rank ≤ n := A.rank_le_card_width.trans $ (fintype.card_fin n).le -lemma rank_mul_le [strong_rank_condition R] (A : matrix m n R) (B : matrix n o R) : +lemma rank_mul_le_left [strong_rank_condition R] (A : matrix m n R) (B : matrix n o R) : (A ⬝ B).rank ≤ A.rank := begin rw [rank, rank, to_lin'_mul], @@ -60,11 +60,28 @@ begin (rank_lt_aleph_0 _ _) (linear_map.rank_comp_le_left _ _), end +include m_fin + +lemma rank_mul_le_right [strong_rank_condition R] (A : matrix l m R) (B : matrix m n R) : + (A ⬝ B).rank ≤ B.rank := +begin + letI := classical.dec_eq m, + rw [rank, rank, to_lin'_mul], + exact finrank_le_finrank_of_rank_le_rank + (linear_map.lift_rank_comp_le_right B.to_lin' A.to_lin') (rank_lt_aleph_0 _ _), +end + +omit m_fin + +lemma rank_mul_le [strong_rank_condition R] (A : matrix m n R) (B : matrix n o R) : + (A ⬝ B).rank ≤ min A.rank B.rank := +le_min (rank_mul_le_left _ _) (rank_mul_le_right _ _) + lemma rank_unit [strong_rank_condition R] (A : (matrix n n R)ˣ) : (A : matrix n n R).rank = fintype.card n := begin refine le_antisymm (rank_le_card_width A) _, - have := rank_mul_le (A : matrix n n R) (↑A⁻¹ : matrix n n R), + have := rank_mul_le_left (A : matrix n n R) (↑A⁻¹ : matrix n n R), rwa [← mul_eq_mul, ← units.coe_mul, mul_inv_self, units.coe_one, rank_one] at this, end diff --git a/src/linear_algebra/dimension.lean b/src/linear_algebra/dimension.lean index 4ba29e9dd770b..f9ac316b7a86a 100644 --- a/src/linear_algebra/dimension.lean +++ b/src/linear_algebra/dimension.lean @@ -1344,10 +1344,27 @@ begin exact linear_map.map_le_range, end +lemma lift_rank_comp_le_right (g : V →ₗ[K] V') (f : V' →ₗ[K] V'') : + cardinal.lift.{v'} (rank (f.comp g)) ≤ cardinal.lift.{v''} (rank g) := +by rw [rank, rank, linear_map.range_comp]; exact lift_rank_map_le _ _ + +/-- The rank of the composition of two maps is less than the minimum of their ranks. -/ +lemma lift_rank_comp_le (g : V →ₗ[K] V') (f : V' →ₗ[K] V'') : + cardinal.lift.{v'} (rank (f.comp g)) ≤ + min (cardinal.lift.{v'} (rank f)) (cardinal.lift.{v''} (rank g)) := +le_min (cardinal.lift_le.mpr $ rank_comp_le_left _ _) (lift_rank_comp_le_right _ _) + variables [add_comm_group V'₁] [module K V'₁] lemma rank_comp_le_right (g : V →ₗ[K] V') (f : V' →ₗ[K] V'₁) : rank (f.comp g) ≤ rank g := -by rw [rank, rank, linear_map.range_comp]; exact rank_map_le _ _ +by simpa only [cardinal.lift_id] using lift_rank_comp_le_right g f + +/-- The rank of the composition of two maps is less than the minimum of their ranks. + +See `lift_rank_comp_le` for the universe-polymorphic version. -/ +lemma rank_comp_le (g : V →ₗ[K] V') (f : V' →ₗ[K] V'₁) : + rank (f.comp g) ≤ min (rank f) (rank g) := +by simpa only [cardinal.lift_id] using lift_rank_comp_le g f end ring From 465d4301d8da5945ef1dc1b29fb34c2f2b315ac4 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 13 Apr 2023 15:25:15 +0000 Subject: [PATCH 0821/1291] chore(linear_algebra/free_module/rank): golf a slow proof (#18804) We already do all the work for this when constructing the basis elsewhere. --- src/linear_algebra/free_module/rank.lean | 10 +++------- 1 file changed, 3 insertions(+), 7 deletions(-) diff --git a/src/linear_algebra/free_module/rank.lean b/src/linear_algebra/free_module/rank.lean index 5fed1512b65cf..875e52ed714a7 100644 --- a/src/linear_algebra/free_module/rank.lean +++ b/src/linear_algebra/free_module/rank.lean @@ -94,13 +94,9 @@ open module.free @[simp] lemma rank_tensor_product : module.rank R (M ⊗[R] N) = lift.{w v} (module.rank R M) * lift.{v w} (module.rank R N) := begin - let ιM := choose_basis_index R M, - let ιN := choose_basis_index R N, - - have h₁ := linear_equiv.lift_rank_eq (tensor_product.congr (repr R M) (repr R N)), - let b : basis (ιM × ιN) R (_ →₀ R) := finsupp.basis_single_one, - rw [linear_equiv.rank_eq (finsupp_tensor_finsupp' R ιM ιN), ← b.mk_eq_rank, mk_prod] at h₁, - rw [lift_inj.1 h₁, rank_eq_card_choose_basis_index R M, rank_eq_card_choose_basis_index R N], + obtain ⟨⟨_, bM⟩⟩ := module.free.exists_basis R M, + obtain ⟨⟨_, bN⟩⟩ := module.free.exists_basis R N, + rw [← bM.mk_eq_rank'', ← bN.mk_eq_rank'', ← (bM.tensor_product bN).mk_eq_rank'', cardinal.mk_prod] end /-- If `M` and `N` lie in the same universe, the rank of `M ⊗[R] N` is From 86add5ce96b35c2cc6ee6946ab458e7302584e21 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 13 Apr 2023 17:32:21 +0000 Subject: [PATCH 0822/1291] refactor(linear_algebra/matrix/rank): remove `decidable_eq` arguments (#18800) `matrix.to_lin' M` is just `matrix.vec_mul_lin M` with an unused decidability argument. We're a bit close to the tide to risk attempting to do a global replacement, so this just: * Refactors some lemmas about `matrix.to_lin'` to be first proven about `matrix.vec_mul_lin` * Changes `matrix.rank` to be defined in terms of the latter. --- src/data/matrix/rank.lean | 31 +++++++------- src/linear_algebra/matrix/to_lin.lean | 58 +++++++++++++++++++++------ 2 files changed, 62 insertions(+), 27 deletions(-) diff --git a/src/data/matrix/rank.lean b/src/data/matrix/rank.lean index db24000f23639..585ec0ad3a580 100644 --- a/src/data/matrix/rank.lean +++ b/src/data/matrix/rank.lean @@ -30,22 +30,23 @@ namespace matrix open finite_dimensional variables {l m n o R : Type*} [m_fin : fintype m] [fintype n] [fintype o] -variables [decidable_eq n] [decidable_eq o] [comm_ring R] +variables [comm_ring R] /-- The rank of a matrix is the rank of its image. -/ -noncomputable def rank (A : matrix m n R) : ℕ := finrank R A.to_lin'.range +noncomputable def rank (A : matrix m n R) : ℕ := finrank R A.mul_vec_lin.range -@[simp] lemma rank_one [strong_rank_condition R] : rank (1 : matrix n n R) = fintype.card n := -by rw [rank, to_lin'_one, linear_map.range_id, finrank_top, finrank_pi] +@[simp] lemma rank_one [strong_rank_condition R] [decidable_eq n] : + rank (1 : matrix n n R) = fintype.card n := +by rw [rank, mul_vec_lin_one, linear_map.range_id, finrank_top, finrank_pi] @[simp] lemma rank_zero [nontrivial R] : rank (0 : matrix m n R) = 0 := -by rw [rank, linear_equiv.map_zero, linear_map.range_zero, finrank_bot] +by rw [rank, mul_vec_lin_zero, linear_map.range_zero, finrank_bot] lemma rank_le_card_width [strong_rank_condition R] (A : matrix m n R) : A.rank ≤ fintype.card n := begin haveI : module.finite R (n → R) := module.finite.pi, haveI : module.free R (n → R) := module.free.pi _ _, - exact A.to_lin'.finrank_range_le.trans_eq (finrank_pi _) + exact A.mul_vec_lin.finrank_range_le.trans_eq (finrank_pi _) end lemma rank_le_width [strong_rank_condition R] {m n : ℕ} (A : matrix (fin m) (fin n) R) : @@ -55,7 +56,7 @@ A.rank_le_card_width.trans $ (fintype.card_fin n).le lemma rank_mul_le_left [strong_rank_condition R] (A : matrix m n R) (B : matrix n o R) : (A ⬝ B).rank ≤ A.rank := begin - rw [rank, rank, to_lin'_mul], + rw [rank, rank, mul_vec_lin_mul], exact cardinal.to_nat_le_of_le_of_lt_aleph_0 (rank_lt_aleph_0 _ _) (linear_map.rank_comp_le_left _ _), end @@ -65,10 +66,9 @@ include m_fin lemma rank_mul_le_right [strong_rank_condition R] (A : matrix l m R) (B : matrix m n R) : (A ⬝ B).rank ≤ B.rank := begin - letI := classical.dec_eq m, - rw [rank, rank, to_lin'_mul], + rw [rank, rank, mul_vec_lin_mul], exact finrank_le_finrank_of_rank_le_rank - (linear_map.lift_rank_comp_le_right B.to_lin' A.to_lin') (rank_lt_aleph_0 _ _), + (linear_map.lift_rank_comp_le_right _ _) (rank_lt_aleph_0 _ _), end omit m_fin @@ -77,7 +77,7 @@ lemma rank_mul_le [strong_rank_condition R] (A : matrix m n R) (B : matrix n o R (A ⬝ B).rank ≤ min A.rank B.rank := le_min (rank_mul_le_left _ _) (rank_mul_le_right _ _) -lemma rank_unit [strong_rank_condition R] (A : (matrix n n R)ˣ) : +lemma rank_unit [strong_rank_condition R] [decidable_eq n] (A : (matrix n n R)ˣ) : (A : matrix n n R).rank = fintype.card n := begin refine le_antisymm (rank_le_card_width A) _, @@ -85,13 +85,14 @@ begin rwa [← mul_eq_mul, ← units.coe_mul, mul_inv_self, units.coe_one, rank_one] at this, end -lemma rank_of_is_unit [strong_rank_condition R] (A : matrix n n R) (h : is_unit A) : +lemma rank_of_is_unit [strong_rank_condition R] [decidable_eq n] (A : matrix n n R) + (h : is_unit A) : A.rank = fintype.card n := by { obtain ⟨A, rfl⟩ := h, exact rank_unit A } include m_fin -lemma rank_eq_finrank_range_to_lin +lemma rank_eq_finrank_range_to_lin [decidable_eq n] {M₁ M₂ : Type*} [add_comm_group M₁] [add_comm_group M₂] [module R M₁] [module R M₂] (A : matrix m n R) (v₁ : basis m R M₁) (v₂ : basis n R M₂) : A.rank = finrank R (to_lin v₂ v₁ A).range := @@ -106,7 +107,7 @@ begin apply linear_map.pi_ext', rintro i, apply linear_map.ext_ring, have aux₁ := to_lin_self (pi.basis_fun R n) (pi.basis_fun R m) A i, have aux₂ := basis.equiv_apply (pi.basis_fun R n) i v₂, - rw [to_lin_eq_to_lin'] at aux₁, + rw [to_lin_eq_to_lin', to_lin'_apply'] at aux₁, rw [pi.basis_fun_apply, linear_map.coe_std_basis] at aux₁ aux₂, simp only [linear_map.comp_apply, e₁, e₂, linear_equiv.coe_coe, equiv.refl_apply, aux₁, aux₂, linear_map.coe_single, to_lin_self, linear_equiv.map_sum, linear_equiv.map_smul, @@ -130,6 +131,6 @@ A.rank_le_card_height.trans $ (fintype.card_fin m).le /-- The rank of a matrix is the rank of the space spanned by its columns. -/ lemma rank_eq_finrank_span_cols (A : matrix m n R) : A.rank = finrank R (submodule.span R (set.range Aᵀ)) := -by rw [rank, matrix.range_to_lin'] +by rw [rank, matrix.range_mul_vec_lin] end matrix diff --git a/src/linear_algebra/matrix/to_lin.lean b/src/linear_algebra/matrix/to_lin.lean index 65524e0897a21..b5f080292e944 100644 --- a/src/linear_algebra/matrix/to_lin.lean +++ b/src/linear_algebra/matrix/to_lin.lean @@ -166,21 +166,54 @@ variables {R : Type*} [comm_semiring R] variables {l m n : Type*} /-- `matrix.mul_vec M` is a linear map. -/ -@[simps] def matrix.mul_vec_lin [fintype n] (M : matrix m n R) : (n → R) →ₗ[R] (m → R) := +def matrix.mul_vec_lin [fintype n] (M : matrix m n R) : (n → R) →ₗ[R] (m → R) := { to_fun := M.mul_vec, map_add' := λ v w, funext (λ i, dot_product_add _ _ _), map_smul' := λ c v, funext (λ i, dot_product_smul _ _ _) } -variables [fintype n] [decidable_eq n] +@[simp] lemma matrix.mul_vec_lin_apply [fintype n] (M : matrix m n R) (v : n → R) : + M.mul_vec_lin v = M.mul_vec v := rfl -lemma matrix.mul_vec_std_basis (M : matrix m n R) (i j) : +@[simp] lemma matrix.mul_vec_lin_zero [fintype n] : matrix.mul_vec_lin (0 : matrix m n R) = 0 := +linear_map.ext zero_mul_vec + +@[simp] lemma matrix.mul_vec_lin_add [fintype n] (M N : matrix m n R) : + (M + N).mul_vec_lin = M.mul_vec_lin + N.mul_vec_lin := +linear_map.ext $ λ _, add_mul_vec _ _ _ + +variables [fintype n] + +@[simp] lemma matrix.mul_vec_lin_one [decidable_eq n] : + matrix.mul_vec_lin (1 : matrix n n R) = id := +by { ext, simp [linear_map.one_apply, std_basis_apply] } + +@[simp] lemma matrix.mul_vec_lin_mul [fintype m] (M : matrix l m R) + (N : matrix m n R) : + matrix.mul_vec_lin (M ⬝ N) = (matrix.mul_vec_lin M).comp (matrix.mul_vec_lin N) := +linear_map.ext $ λ x, (mul_vec_mul_vec _ _ _).symm + +lemma matrix.ker_mul_vec_lin_eq_bot_iff {M : matrix n n R} : + M.mul_vec_lin.ker = ⊥ ↔ ∀ v, M.mul_vec v = 0 → v = 0 := +by simp only [submodule.eq_bot_iff, linear_map.mem_ker, matrix.mul_vec_lin_apply] + +lemma matrix.mul_vec_std_basis [decidable_eq n] (M : matrix m n R) (i j) : M.mul_vec (std_basis R (λ _, R) j 1) i = M i j := (congr_fun (matrix.mul_vec_single _ _ (1 : R)) i).trans $ mul_one _ -@[simp] lemma matrix.mul_vec_std_basis_apply (M : matrix m n R) (j) : +@[simp] lemma matrix.mul_vec_std_basis_apply [decidable_eq n] (M : matrix m n R) (j) : M.mul_vec (std_basis R (λ _, R) j 1) = Mᵀ j := funext $ λ i, matrix.mul_vec_std_basis M i j +lemma matrix.range_mul_vec_lin (M : matrix m n R) : M.mul_vec_lin.range = span R (range Mᵀ) := +begin + letI := classical.dec_eq n, + simp_rw [range_eq_map, ←supr_range_std_basis, submodule.map_supr, range_eq_map, + ←ideal.span_singleton_one, ideal.span, submodule.map_span, image_image, image_singleton, + matrix.mul_vec_lin_apply, M.mul_vec_std_basis_apply, supr_span, range_eq_Union] +end + +variables [decidable_eq n] + /-- Linear maps `(n → R) →ₗ[R] (m → R)` are linearly equivalent to `matrix m n R`. -/ def linear_map.to_matrix' : ((n → R) →ₗ[R] (m → R)) ≃ₗ[R] matrix m n R := { to_fun := λ f, of (λ i j, f (std_basis R (λ _, R) j 1) i), @@ -197,10 +230,14 @@ def linear_map.to_matrix' : ((n → R) →ₗ[R] (m → R)) ≃ₗ[R] matrix m n map_smul' := λ c f, by { ext i j, simp only [pi.smul_apply, linear_map.smul_apply, ring_hom.id_apply, of_apply] } } -/-- A `matrix m n R` is linearly equivalent to a linear map `(n → R) →ₗ[R] (m → R)`. -/ +/-- A `matrix m n R` is linearly equivalent to a linear map `(n → R) →ₗ[R] (m → R)`. + +Note that the forward-direction does not require `decidable_eq` and is `matrix.vec_mul_lin`. -/ def matrix.to_lin' : matrix m n R ≃ₗ[R] ((n → R) →ₗ[R] (m → R)) := linear_map.to_matrix'.symm +lemma matrix.to_lin'_apply' (M : matrix m n R) : matrix.to_lin' M = M.mul_vec_lin := rfl + @[simp] lemma linear_map.to_matrix'_symm : (linear_map.to_matrix'.symm : matrix m n R ≃ₗ[R] _) = matrix.to_lin' := rfl @@ -232,8 +269,7 @@ end matrix.to_lin' M v = M.mul_vec v := rfl @[simp] lemma matrix.to_lin'_one : - matrix.to_lin' (1 : matrix n n R) = id := -by { ext, simp [linear_map.one_apply, std_basis_apply] } + matrix.to_lin' (1 : matrix n n R) = id := matrix.mul_vec_lin_one @[simp] lemma linear_map.to_matrix'_id : (linear_map.to_matrix' (linear_map.id : (n → R) →ₗ[R] (n → R))) = 1 := @@ -241,7 +277,7 @@ by { ext, rw [matrix.one_apply, linear_map.to_matrix'_apply, id_apply] } @[simp] lemma matrix.to_lin'_mul [fintype m] [decidable_eq m] (M : matrix l m R) (N : matrix m n R) : matrix.to_lin' (M ⬝ N) = (matrix.to_lin' M).comp (matrix.to_lin' N) := -linear_map.ext $ λ x, (mul_vec_mul_vec _ _ _).symm +matrix.mul_vec_lin_mul _ _ /-- Shortcut lemma for `matrix.to_lin'_mul` and `linear_map.comp_apply` -/ lemma matrix.to_lin'_mul_apply [fintype m] [decidable_eq m] (M : matrix l m R) @@ -266,12 +302,10 @@ by simp [module.algebra_map_End_eq_smul_id] lemma matrix.ker_to_lin'_eq_bot_iff {M : matrix n n R} : M.to_lin'.ker = ⊥ ↔ ∀ v, M.mul_vec v = 0 → v = 0 := -by simp only [submodule.eq_bot_iff, linear_map.mem_ker, matrix.to_lin'_apply] +matrix.ker_mul_vec_lin_eq_bot_iff lemma matrix.range_to_lin' (M : matrix m n R) : M.to_lin'.range = span R (range Mᵀ) := -by simp_rw [range_eq_map, ←supr_range_std_basis, submodule.map_supr, range_eq_map, - ←ideal.span_singleton_one, ideal.span, submodule.map_span, image_image, image_singleton, - matrix.to_lin'_apply, M.mul_vec_std_basis_apply, supr_span, range_eq_Union] +matrix.range_mul_vec_lin _ /-- If `M` and `M'` are each other's inverse matrices, they provide an equivalence between `m → A` and `n → A` corresponding to `M.mul_vec` and `M'.mul_vec`. -/ From 7ebf83ed9c262adbf983ef64d5e8c2ae94b625f4 Mon Sep 17 00:00:00 2001 From: Moritz Doll Date: Thu, 13 Apr 2023 19:25:22 +0000 Subject: [PATCH 0823/1291] chore(analysis/seminorm): add new le_def/lt_def and renaming (#18801) Co-authored-by: Eric Wieser --- src/analysis/seminorm.lean | 9 ++++++--- 1 file changed, 6 insertions(+), 3 deletions(-) diff --git a/src/analysis/seminorm.lean b/src/analysis/seminorm.lean index 450906a1b3a68..ded9d63911138 100644 --- a/src/analysis/seminorm.lean +++ b/src/analysis/seminorm.lean @@ -213,8 +213,11 @@ ext $ λ x, real.smul_max _ _ instance : partial_order (seminorm 𝕜 E) := partial_order.lift _ fun_like.coe_injective -lemma le_def (p q : seminorm 𝕜 E) : p ≤ q ↔ (p : E → ℝ) ≤ q := iff.rfl -lemma lt_def (p q : seminorm 𝕜 E) : p < q ↔ (p : E → ℝ) < q := iff.rfl +@[simp, norm_cast] lemma coe_le_coe {p q : seminorm 𝕜 E} : (p : E → ℝ) ≤ q ↔ p ≤ q := iff.rfl +@[simp, norm_cast] lemma coe_lt_coe {p q : seminorm 𝕜 E} : (p : E → ℝ) < q ↔ p < q := iff.rfl + +lemma le_def {p q : seminorm 𝕜 E} : p ≤ q ↔ ∀ x, p x ≤ q x := iff.rfl +lemma lt_def {p q : seminorm 𝕜 E} : p < q ↔ p ≤ q ∧ ∃ x, p x < q x := pi.lt_def instance : semilattice_sup (seminorm 𝕜 E) := function.injective.semilattice_sup _ fun_like.coe_injective coe_sup @@ -285,7 +288,7 @@ lemma bot_eq_zero : (⊥ : seminorm 𝕜 E) = 0 := rfl lemma smul_le_smul {p q : seminorm 𝕜 E} {a b : ℝ≥0} (hpq : p ≤ q) (hab : a ≤ b) : a • p ≤ b • q := begin - simp_rw [le_def, pi.le_def, coe_smul], + simp_rw [le_def, coe_smul], intros x, simp_rw [pi.smul_apply, nnreal.smul_def, smul_eq_mul], exact mul_le_mul hab (hpq x) (map_nonneg p x) (nnreal.coe_nonneg b), From 820b22968a2bc4a47ce5cf1d2f36a9ebe52510aa Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 13 Apr 2023 21:12:11 +0000 Subject: [PATCH 0824/1291] feat(data/matrix/reflection): add `mul_vec` and `vec_mul` (#18805) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This follows the pattern already established by `mul`. The motivation was an example on Zulip which wanted to compute the product of ```lean def M := !![(2:ℂ), 0, 0; 0, 1, 0; 0, 0, 1] def v := ![(0:ℂ), 0, 1] ``` As before, the meta code that makes this pleasant to use is absent, but I will add it along with the rest of the meta code in #15738. --- src/data/matrix/reflection.lean | 65 ++++++++++++++++++++++++++++++--- 1 file changed, 60 insertions(+), 5 deletions(-) diff --git a/src/data/matrix/reflection.lean b/src/data/matrix/reflection.lean index ea0f5921db200..dd9c6d06a48ec 100644 --- a/src/data/matrix/reflection.lean +++ b/src/data/matrix/reflection.lean @@ -25,6 +25,8 @@ corresponding `*_eq` lemmas to be used in a place where they are definitionally * `matrix.transposeᵣ` * `matrix.dot_productᵣ` * `matrix.mulᵣ` +* `matrix.mul_vecᵣ` +* `matrix.vec_mulᵣ` * `matrix.eta_expand` -/ @@ -130,17 +132,17 @@ example (a b c d : α) [has_mul α] [add_comm_monoid α] : def mulᵣ [has_mul α] [has_add α] [has_zero α] (A : matrix (fin l) (fin m) α) (B : matrix (fin m) (fin n) α) : matrix (fin l) (fin n) α := -of $ fin_vec.map (λ v₁, fin_vec.map (λ v₂, dot_productᵣ v₁ v₂) (transposeᵣ B)) A +of $ fin_vec.map (λ v₁, fin_vec.map (λ v₂, dot_productᵣ v₁ v₂) Bᵀ) A /-- This can be used to prove ```lean -example [add_comm_monoid α] [has_mul α] - (a₁₁ a₁₂ a₁₃ a₂₁ a₂₂ a₂₃ a₃₁ a₃₂ a₃₃ b₁₁ b₁₂ b₁₃ b₂₁ b₂₂ b₂₃ b₃₁ b₃₂ b₃₃ : α) : +example [add_comm_monoid α] [has_mul α] (a₁₁ a₁₂ a₂₁ a₂₂ b₁₁ b₁₂ b₂₁ b₂₂ : α) : !![a₁₁, a₁₂; a₂₁, a₂₂] ⬝ !![b₁₁, b₁₂; b₂₁, b₂₂] = !![a₁₁*b₁₁ + a₁₂*b₂₁, a₁₁*b₁₂ + a₁₂*b₂₂; a₂₁*b₁₁ + a₂₂*b₂₁, a₂₁*b₁₂ + a₂₂*b₂₂] := +(mulᵣ_eq _ _).symm ``` -/ @[simp] @@ -152,8 +154,7 @@ begin refl, end -example [add_comm_monoid α] [has_mul α] - (a₁₁ a₁₂ a₁₃ a₂₁ a₂₂ a₂₃ a₃₁ a₃₂ a₃₃ b₁₁ b₁₂ b₁₃ b₂₁ b₂₂ b₂₃ b₃₁ b₃₂ b₃₃ : α) : +example [add_comm_monoid α] [has_mul α] (a₁₁ a₁₂ a₂₁ a₂₂ b₁₁ b₁₂ b₂₁ b₂₂ : α) : !![a₁₁, a₁₂; a₂₁, a₂₂].mul !![b₁₁, b₁₂; b₂₁, b₂₂] = @@ -161,6 +162,60 @@ example [add_comm_monoid α] [has_mul α] a₂₁*b₁₁ + a₂₂*b₂₁, a₂₁*b₁₂ + a₂₂*b₂₂] := (mulᵣ_eq _ _).symm +/-- `matrix.mul_vec` with better defeq for `fin` -/ +def mul_vecᵣ [has_mul α] [has_add α] [has_zero α] (A : matrix (fin l) (fin m) α) (v : fin m → α) : + fin l → α := +fin_vec.map (λ a, dot_productᵣ a v) A + +/-- This can be used to prove +```lean +example [non_unital_non_assoc_semiring α] (a₁₁ a₁₂ a₂₁ a₂₂ b₁ b₂ : α) : + !![a₁₁, a₁₂; + a₂₁, a₂₂].mul_vec ![b₁, b₂] = ![a₁₁*b₁ + a₁₂*b₂, a₂₁*b₁ + a₂₂*b₂] := +(mul_vecᵣ_eq _ _).symm +``` +-/ +@[simp] +lemma mul_vecᵣ_eq [non_unital_non_assoc_semiring α] + (A : matrix (fin l) (fin m) α) (v : fin m → α) : + mul_vecᵣ A v = A.mul_vec v := +begin + simp [mul_vecᵣ, function.comp], + refl, +end + +example [non_unital_non_assoc_semiring α] (a₁₁ a₁₂ a₂₁ a₂₂ b₁ b₂ : α) : + !![a₁₁, a₁₂; + a₂₁, a₂₂].mul_vec ![b₁, b₂] = ![a₁₁*b₁ + a₁₂*b₂, a₂₁*b₁ + a₂₂*b₂] := +(mul_vecᵣ_eq _ _).symm + +/-- `matrix.vec_mul` with better defeq for `fin` -/ +def vec_mulᵣ [has_mul α] [has_add α] [has_zero α] (v : fin l → α) (A : matrix (fin l) (fin m) α): + fin m → α := +fin_vec.map (λ a, dot_productᵣ v a) Aᵀ + +/-- This can be used to prove +```lean +example [non_unital_non_assoc_semiring α] (a₁₁ a₁₂ a₂₁ a₂₂ b₁ b₂ : α) : + vec_mul ![b₁, b₂] !![a₁₁, a₁₂; + a₂₁, a₂₂] = ![b₁*a₁₁ + b₂*a₂₁, b₁*a₁₂ + b₂*a₂₂] := +(vec_mulᵣ_eq _ _).symm +``` +-/ +@[simp] +lemma vec_mulᵣ_eq [non_unital_non_assoc_semiring α] + (v : fin l → α) (A : matrix (fin l) (fin m) α) : + vec_mulᵣ v A = vec_mul v A := +begin + simp [vec_mulᵣ, function.comp], + refl, +end + +example [non_unital_non_assoc_semiring α] (a₁₁ a₁₂ a₂₁ a₂₂ b₁ b₂ : α) : + vec_mul ![b₁, b₂] !![a₁₁, a₁₂; + a₂₁, a₂₂] = ![b₁*a₁₁ + b₂*a₂₁, b₁*a₁₂ + b₂*a₂₂] := +(vec_mulᵣ_eq _ _).symm + /-- Expand `A` to `!![A 0 0, ...; ..., A m n]` -/ def eta_expand {m n} (A : matrix (fin m) (fin n) α) : matrix (fin m) (fin n) α := matrix.of (fin_vec.eta_expand (λ i, fin_vec.eta_expand (λ j, A i j))) From 2cf3ee29d9361ee308a14a46d820a49b8856d041 Mon Sep 17 00:00:00 2001 From: Moritz Doll Date: Fri, 14 Apr 2023 06:11:40 +0000 Subject: [PATCH 0825/1291] =?UTF-8?q?feat(analysis/calculus/cont=5Fdiff=5F?= =?UTF-8?q?def):=20no=20continuity=20is=20needed=20for=20has=5Fftaylor=5Fs?= =?UTF-8?q?eries=5Fup=5Fto=20if=20n=20=3D=20=E2=88=9E=20(#18808)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We also add `has_ftaylor_series_up_to_iff_top` in order that the notation is consistent between `has_ftaylor_series_up_to` and `has_ftaylor_series_up_to_on`. --- src/analysis/calculus/cont_diff_def.lean | 22 ++++++++++++++++++++++ 1 file changed, 22 insertions(+) diff --git a/src/analysis/calculus/cont_diff_def.lean b/src/analysis/calculus/cont_diff_def.lean index 2d4bf3367fc70..8720cd1832c86 100644 --- a/src/analysis/calculus/cont_diff_def.lean +++ b/src/analysis/calculus/cont_diff_def.lean @@ -248,6 +248,16 @@ begin apply (H m).cont m le_rfl } } end +/-- In the case that `n = ∞` we don't need the continuity assumption in +`has_ftaylor_series_up_to_on`. -/ +lemma has_ftaylor_series_up_to_on_top_iff' : has_ftaylor_series_up_to_on ∞ f p s ↔ + (∀ x ∈ s, (p x 0).uncurry0 = f x) ∧ + (∀ (m : ℕ), ∀ x ∈ s, has_fderiv_within_at (λ y, p y m) (p x m.succ).curry_left s x) := +-- Everything except for the continuity is trivial: +⟨λ h, ⟨h.1, λ m, h.2 m (with_top.coe_lt_top m)⟩, λ h, ⟨h.1, λ m _, h.2 m, λ m _ x hx, + -- The continuity follows from the existence of a derivative: + (h.2 m x hx).continuous_within_at⟩⟩ + /-- If a function has a Taylor series at order at least `1`, then the term of order `1` of this series is a derivative of `f`. -/ lemma has_ftaylor_series_up_to_on.has_fderiv_within_at @@ -1234,6 +1244,18 @@ lemma has_ftaylor_series_up_to_zero_iff : by simp [has_ftaylor_series_up_to_on_univ_iff.symm, continuous_iff_continuous_on_univ, has_ftaylor_series_up_to_on_zero_iff] +lemma has_ftaylor_series_up_to_top_iff : has_ftaylor_series_up_to ∞ f p ↔ + ∀ (n : ℕ), has_ftaylor_series_up_to n f p := +by simp only [← has_ftaylor_series_up_to_on_univ_iff, has_ftaylor_series_up_to_on_top_iff] + +/-- In the case that `n = ∞` we don't need the continuity assumption in +`has_ftaylor_series_up_to`. -/ +lemma has_ftaylor_series_up_to_top_iff' : has_ftaylor_series_up_to ∞ f p ↔ + (∀ x, (p x 0).uncurry0 = f x) ∧ + (∀ (m : ℕ) x, has_fderiv_at (λ y, p y m) (p x m.succ).curry_left x) := +by simp only [← has_ftaylor_series_up_to_on_univ_iff, has_ftaylor_series_up_to_on_top_iff', + mem_univ, forall_true_left, has_fderiv_within_at_univ] + /-- If a function has a Taylor series at order at least `1`, then the term of order `1` of this series is a derivative of `f`. -/ lemma has_ftaylor_series_up_to.has_fderiv_at From 0e2aab2b0d521f060f62a14d2cf2e2c54e8491d6 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 14 Apr 2023 17:40:35 +0000 Subject: [PATCH 0826/1291] feat(data/matrix/basic): more lemmas about submatrix (#18738) This adds trivial results about reordering rows and columns in a matrix: * Reordering can be moved around `dot_product`, `mul_vec`, and `vec_mul` (previouly we only had `mul`) * Reordering can be moved through `adjugate` * Reordering can be moved through row and column updates. * Reordering can be moved through `to_lin'` * Reordering does not affect `matrix.rank` Also adds some missing `of` wrappers. Lemmas about reindexing are useful when working with block matrices, as they make it possible to make symmetry arguments. --- src/data/matrix/basic.lean | 75 ++++++++++++++++++++- src/data/matrix/rank.lean | 21 ++++++ src/linear_algebra/matrix/adjugate.lean | 33 ++++++++- src/linear_algebra/matrix/to_lin.lean | 26 ++++++- src/linear_algebra/matrix/transvection.lean | 12 ++-- 5 files changed, 155 insertions(+), 12 deletions(-) diff --git a/src/data/matrix/basic.lean b/src/data/matrix/basic.lean index 65388afef5547..44e0d83386a43 100644 --- a/src/data/matrix/basic.lean +++ b/src/data/matrix/basic.lean @@ -493,6 +493,19 @@ by simp [dot_product, mul_add, finset.sum_add_distrib] (sum.elim u x) ⬝ᵥ (sum.elim v y) = u ⬝ᵥ v + x ⬝ᵥ y := by simp [dot_product] +/-- Permuting a vector on the left of a dot product can be transferred to the right. -/ +@[simp] lemma comp_equiv_symm_dot_product (e : m ≃ n) : (u ∘ e.symm) ⬝ᵥ x = u ⬝ᵥ (x ∘ e) := +(e.sum_comp _).symm.trans $ finset.sum_congr rfl $ λ _ _, + by simp only [function.comp, equiv.symm_apply_apply] + +/-- Permuting a vector on the right of a dot product can be transferred to the left. -/ +@[simp] lemma dot_product_comp_equiv_symm (e : n ≃ m) : u ⬝ᵥ (x ∘ e.symm) = (u ∘ e) ⬝ᵥ x := +by simpa only [equiv.symm_symm] using (comp_equiv_symm_dot_product u x e.symm).symm + +/-- Permuting vectors on both sides of a dot product is a no-op. -/ +@[simp] lemma comp_equiv_dot_product_comp_equiv (e : m ≃ n) : (x ∘ e) ⬝ᵥ (y ∘ e) = x ⬝ᵥ y := +by simp only [←dot_product_comp_equiv_symm, function.comp, equiv.apply_symm_apply] + end non_unital_non_assoc_semiring section non_unital_non_assoc_semiring_decidable @@ -1793,6 +1806,16 @@ lemma submatrix_mul_equiv [fintype n] [fintype o] [add_comm_monoid α] [has_mul (M.submatrix e₁ e₂) ⬝ (N.submatrix e₂ e₃) = (M ⬝ N).submatrix e₁ e₃ := (submatrix_mul M N e₁ e₂ e₃ e₂.bijective).symm +lemma submatrix_mul_vec_equiv [fintype n] [fintype o] [non_unital_non_assoc_semiring α] + (M : matrix m n α) (v : o → α) (e₁ : l → m) (e₂ : o ≃ n) : + (M.submatrix e₁ e₂).mul_vec v = M.mul_vec (v ∘ e₂.symm) ∘ e₁ := +funext $ λ i, eq.symm (dot_product_comp_equiv_symm _ _ _) + +lemma submatrix_vec_mul_equiv [fintype l] [fintype m] [non_unital_non_assoc_semiring α] + (M : matrix m n α) (v : l → α) (e₁ : l ≃ m) (e₂ : o → n) : + vec_mul v (M.submatrix e₁ e₂) = vec_mul (v ∘ e₁.symm) M ∘ e₂ := +funext $ λ i, eq.symm (comp_equiv_symm_dot_product _ _ _) + lemma mul_submatrix_one [fintype n] [fintype o] [non_assoc_semiring α] [decidable_eq o] (e₁ : n ≃ o) (e₂ : l → o) (M : matrix m n α) : M ⬝ (1 : matrix o o α).submatrix e₁ e₂ = submatrix M id (e₁.symm ∘ e₂) := @@ -1947,11 +1970,11 @@ section update /-- Update, i.e. replace the `i`th row of matrix `A` with the values in `b`. -/ def update_row [decidable_eq m] (M : matrix m n α) (i : m) (b : n → α) : matrix m n α := -function.update M i b +of $ function.update M i b /-- Update, i.e. replace the `j`th column of matrix `A` with the values in `b`. -/ def update_column [decidable_eq n] (M : matrix m n α) (j : n) (b : m → α) : matrix m n α := -λ i, function.update (M i) j (b i) +of $ λ i, function.update (M i) j (b i) variables {M : matrix m n α} {i : m} {j : n} {b : n → α} {c : m → α} @@ -2074,6 +2097,54 @@ lemma diagonal_update_row_single [decidable_eq n] [has_zero α] (v : n → α) ( (diagonal v).update_row i (pi.single i x) = diagonal (function.update v i x) := by rw [←diagonal_transpose, update_row_transpose, diagonal_update_column_single, diagonal_transpose] +/-! Updating rows and columns commutes in the obvious way with reindexing the matrix. -/ + +lemma update_row_submatrix_equiv [decidable_eq l] [decidable_eq m] + (A : matrix m n α) (i : l) (r : o → α) (e : l ≃ m) (f : o ≃ n) : + update_row (A.submatrix e f) i r = (A.update_row (e i) (λ j, r (f.symm j))).submatrix e f := +begin + ext i' j, + simp only [submatrix_apply, update_row_apply, equiv.apply_eq_iff_eq, equiv.symm_apply_apply], +end + +lemma submatrix_update_row_equiv [decidable_eq l] [decidable_eq m] + (A : matrix m n α) (i : m) (r : n → α) (e : l ≃ m) (f : o ≃ n) : + (A.update_row i r).submatrix e f = update_row (A.submatrix e f) (e.symm i) (λ i, r (f i)) := +eq.trans (by simp_rw equiv.apply_symm_apply) (update_row_submatrix_equiv A _ _ e f).symm + +lemma update_column_submatrix_equiv [decidable_eq o] [decidable_eq n] + (A : matrix m n α) (j : o) (c : l → α) (e : l ≃ m) (f : o ≃ n) : + update_column (A.submatrix e f) j c = (A.update_column (f j) (λ i, c (e.symm i))).submatrix e f := +by simpa only [←transpose_submatrix, update_row_transpose] using + congr_arg transpose (update_row_submatrix_equiv Aᵀ j c f e) + +lemma submatrix_update_column_equiv [decidable_eq o] [decidable_eq n] + (A : matrix m n α) (j : n) (c : m → α) (e : l ≃ m) (f : o ≃ n) : + (A.update_column j c).submatrix e f = update_column (A.submatrix e f) (f.symm j) (λ i, c (e i)) := +eq.trans (by simp_rw equiv.apply_symm_apply) (update_column_submatrix_equiv A _ _ e f).symm + +/-! `reindex` versions of the above `submatrix` lemmas for convenience. -/ + +lemma update_row_reindex [decidable_eq l] [decidable_eq m] + (A : matrix m n α) (i : l) (r : o → α) (e : m ≃ l) (f : n ≃ o) : + update_row (reindex e f A) i r = reindex e f (A.update_row (e.symm i) (λ j, r (f j))) := +update_row_submatrix_equiv _ _ _ _ _ + +lemma reindex_update_row [decidable_eq l] [decidable_eq m] + (A : matrix m n α) (i : m) (r : n → α) (e : m ≃ l) (f : n ≃ o) : + reindex e f (A.update_row i r) = update_row (reindex e f A) (e i) (λ i, r (f.symm i)) := +submatrix_update_row_equiv _ _ _ _ _ + +lemma update_column_reindex [decidable_eq o] [decidable_eq n] + (A : matrix m n α) (j : o) (c : l → α) (e : m ≃ l) (f : n ≃ o) : + update_column (reindex e f A) j c = reindex e f (A.update_column (f.symm j) (λ i, c (e i))) := +update_column_submatrix_equiv _ _ _ _ _ + +lemma reindex_update_column [decidable_eq o] [decidable_eq n] + (A : matrix m n α) (j : n) (c : m → α) (e : m ≃ l) (f : n ≃ o) : + reindex e f (A.update_column j c) = update_column (reindex e f A) (f j) (λ i, c (e.symm i)) := +submatrix_update_column_equiv _ _ _ _ _ + end update end matrix diff --git a/src/data/matrix/rank.lean b/src/data/matrix/rank.lean index 585ec0ad3a580..1e491bb9adbfe 100644 --- a/src/data/matrix/rank.lean +++ b/src/data/matrix/rank.lean @@ -90,6 +90,27 @@ lemma rank_of_is_unit [strong_rank_condition R] [decidable_eq n] (A : matrix n n A.rank = fintype.card n := by { obtain ⟨A, rfl⟩ := h, exact rank_unit A } +/-- Taking a subset of the rows and permuting the columns reduces the rank. -/ +lemma rank_submatrix_le [strong_rank_condition R] [fintype m] (f : n → m) (e : n ≃ m) + (A : matrix m m R) : + rank (A.submatrix f e) ≤ rank A := +begin + rw [rank, rank, mul_vec_lin_submatrix, linear_map.range_comp, linear_map.range_comp, + (show linear_map.fun_left R R e.symm = linear_equiv.fun_congr_left R R e.symm, from rfl), + linear_equiv.range, submodule.map_top], + -- TODO: generalize `finite_dimensional.finrank_map_le` and use it here + exact finrank_le_finrank_of_rank_le_rank (lift_rank_map_le _ _) (rank_lt_aleph_0 _ _), +end + +lemma rank_reindex [fintype m] (e₁ e₂ : m ≃ n) (A : matrix m m R) : + rank (reindex e₁ e₂ A) = rank A := +by rw [rank, rank, mul_vec_lin_reindex, linear_map.range_comp, linear_map.range_comp, + linear_equiv.range, submodule.map_top, linear_equiv.finrank_map_eq] + +@[simp] lemma rank_submatrix [fintype m] (A : matrix m m R) (e₁ e₂ : n ≃ m) : + rank (A.submatrix e₁ e₂) = rank A := +by simpa only [reindex_apply] using rank_reindex e₁.symm e₂.symm A + include m_fin lemma rank_eq_finrank_range_to_lin [decidable_eq n] diff --git a/src/linear_algebra/matrix/adjugate.lean b/src/linear_algebra/matrix/adjugate.lean index ba7511931f031..a3868b6d8d3f7 100644 --- a/src/linear_algebra/matrix/adjugate.lean +++ b/src/linear_algebra/matrix/adjugate.lean @@ -40,8 +40,9 @@ cramer, cramer's rule, adjugate -/ namespace matrix -universes u v -variables {n : Type u} [decidable_eq n] [fintype n] {α : Type v} [comm_ring α] +universes u v w +variables {m : Type u} {n : Type v} {α : Type w} +variables [decidable_eq n] [fintype n] [decidable_eq m] [fintype m] [comm_ring α] open_locale matrix big_operators polynomial open equiv equiv.perm finset @@ -98,7 +99,7 @@ begin split_ifs with h, { -- i = j: this entry should be `A.det` subst h, - simp only [update_column_transpose, det_transpose, update_row, function.update_eq_self] }, + simp only [update_column_transpose, det_transpose, update_row_eq_self] }, { -- i ≠ j: this entry should be 0 rw [update_column_transpose, det_transpose], apply det_zero_of_row_eq h, @@ -151,6 +152,18 @@ calc ∑ x in s, cramer A (λ j, f j x) i ... = cramer A (λ (j : n), ∑ x in s, f j x) i : by { rw [sum_cramer, cramer_apply], congr' with j, apply finset.sum_apply } +lemma cramer_submatrix_equiv (A : matrix m m α) (e : n ≃ m) (b : n → α) : + cramer (A.submatrix e e) b = cramer A (b ∘ e.symm) ∘ e := +begin + ext i, + simp_rw [function.comp_apply, cramer_apply, update_column_submatrix_equiv, + det_submatrix_equiv_self e], +end + +lemma cramer_reindex (e : m ≃ n) (A : matrix m m α) (b : n → α) : + cramer (reindex e e A) b = cramer A (b ∘ e) ∘ e.symm := +cramer_submatrix_equiv _ _ _ + end cramer section adjugate @@ -210,6 +223,20 @@ begin exact h ((symm_apply_eq σ).mp h') } end +@[simp] lemma adjugate_submatrix_equiv_self (e : n ≃ m) (A : matrix m m α) : + adjugate (A.submatrix e e) = (adjugate A).submatrix e e := +begin + ext i j, + rw [adjugate_apply, submatrix_apply, adjugate_apply, ← det_submatrix_equiv_self e, + update_row_submatrix_equiv], + congr, + exact function.update_comp_equiv _ e.symm _ _, +end + +lemma adjugate_reindex (e : m ≃ n) (A : matrix m m α) : + adjugate (reindex e e A) = reindex e e (adjugate A) := +adjugate_submatrix_equiv_self _ _ + /-- Since the map `b ↦ cramer A b` is linear in `b`, it must be multiplication by some matrix. This matrix is `A.adjugate`. -/ lemma cramer_eq_adjugate_mul_vec (A : matrix n n α) (b : n → α) : diff --git a/src/linear_algebra/matrix/to_lin.lean b/src/linear_algebra/matrix/to_lin.lean index b5f080292e944..7e5817ef7353a 100644 --- a/src/linear_algebra/matrix/to_lin.lean +++ b/src/linear_algebra/matrix/to_lin.lean @@ -163,7 +163,7 @@ This should eventually be remedied. section to_matrix' variables {R : Type*} [comm_semiring R] -variables {l m n : Type*} +variables {k l m n : Type*} /-- `matrix.mul_vec M` is a linear map. -/ def matrix.mul_vec_lin [fintype n] (M : matrix m n R) : (n → R) →ₗ[R] (m → R) := @@ -181,6 +181,18 @@ linear_map.ext zero_mul_vec (M + N).mul_vec_lin = M.mul_vec_lin + N.mul_vec_lin := linear_map.ext $ λ _, add_mul_vec _ _ _ +lemma matrix.mul_vec_lin_submatrix [fintype n] [fintype l] (f₁ : m → k) (e₂ : n ≃ l) + (M : matrix k l R) : + (M.submatrix f₁ e₂).mul_vec_lin = fun_left R R f₁ ∘ₗ M.mul_vec_lin ∘ₗ fun_left _ _ e₂.symm := +linear_map.ext $ λ x, submatrix_mul_vec_equiv _ _ _ _ + +/-- A variant of `matrix.mul_vec_lin_submatrix` that keeps around `linear_equiv`s. -/ +lemma matrix.mul_vec_lin_reindex [fintype n] [fintype l] (e₁ : k ≃ m) (e₂ : l ≃ n) + (M : matrix k l R) : + (reindex e₁ e₂ M).mul_vec_lin = ↑(linear_equiv.fun_congr_left R R e₁.symm) + ∘ₗ M.mul_vec_lin ∘ₗ ↑(linear_equiv.fun_congr_left R R e₂) := +matrix.mul_vec_lin_submatrix _ _ _ + variables [fintype n] @[simp] lemma matrix.mul_vec_lin_one [decidable_eq n] : @@ -279,6 +291,18 @@ by { ext, rw [matrix.one_apply, linear_map.to_matrix'_apply, id_apply] } (N : matrix m n R) : matrix.to_lin' (M ⬝ N) = (matrix.to_lin' M).comp (matrix.to_lin' N) := matrix.mul_vec_lin_mul _ _ +@[simp] lemma matrix.to_lin'_submatrix [fintype l] [decidable_eq l] (f₁ : m → k) (e₂ : n ≃ l) + (M : matrix k l R) : + (M.submatrix f₁ e₂).to_lin' = fun_left R R f₁ ∘ₗ M.to_lin' ∘ₗ fun_left _ _ e₂.symm := +matrix.mul_vec_lin_submatrix _ _ _ + +/-- A variant of `matrix.to_lin'_submatrix` that keeps around `linear_equiv`s. -/ +lemma matrix.to_lin'_reindex [fintype l] [decidable_eq l] (e₁ : k ≃ m) (e₂ : l ≃ n) + (M : matrix k l R) : + (reindex e₁ e₂ M).to_lin' = ↑(linear_equiv.fun_congr_left R R e₁.symm) + ∘ₗ M.to_lin' ∘ₗ ↑(linear_equiv.fun_congr_left R R e₂) := +matrix.mul_vec_lin_reindex _ _ _ + /-- Shortcut lemma for `matrix.to_lin'_mul` and `linear_map.comp_apply` -/ lemma matrix.to_lin'_mul_apply [fintype m] [decidable_eq m] (M : matrix l m R) (N : matrix m n R) (x) : matrix.to_lin' (M ⬝ N) x = (matrix.to_lin' M (matrix.to_lin' N x)) := diff --git a/src/linear_algebra/matrix/transvection.lean b/src/linear_algebra/matrix/transvection.lean index 3d7da8b6dd158..e702144be94ef 100644 --- a/src/linear_algebra/matrix/transvection.lean +++ b/src/linear_algebra/matrix/transvection.lean @@ -90,13 +90,13 @@ begin casesI nonempty_fintype n, ext a b, by_cases ha : i = a, by_cases hb : j = b, - { simp only [update_row, transvection, ha, hb, function.update_same, std_basis_matrix.apply_same, - pi.add_apply, one_apply_eq, pi.smul_apply, mul_one, algebra.id.smul_eq_mul], }, - { simp only [update_row, transvection, ha, hb, std_basis_matrix.apply_of_ne, function.update_same, - pi.add_apply, ne.def, not_false_iff, pi.smul_apply, and_false, one_apply_ne, + { simp only [update_row_self, transvection, ha, hb, pi.add_apply, + std_basis_matrix.apply_same, one_apply_eq, pi.smul_apply, mul_one, algebra.id.smul_eq_mul], }, + { simp only [update_row_self, transvection, ha, hb, std_basis_matrix.apply_of_ne, pi.add_apply, + ne.def, not_false_iff, pi.smul_apply, and_false, one_apply_ne, algebra.id.smul_eq_mul, mul_zero] }, - { simp only [update_row, transvection, ha, ne.symm ha, std_basis_matrix.apply_of_ne, add_zero, - algebra.id.smul_eq_mul, function.update_noteq, ne.def, not_false_iff, dmatrix.add_apply, + { simp only [update_row_ne, transvection, ha, ne.symm ha, std_basis_matrix.apply_of_ne, add_zero, + algebra.id.smul_eq_mul, ne.def, not_false_iff, dmatrix.add_apply, pi.smul_apply, mul_zero, false_and] }, end From de29c328903507bb7aff506af9135f4bdaf1849c Mon Sep 17 00:00:00 2001 From: Kevin Buzzard Date: Fri, 14 Apr 2023 19:47:13 +0000 Subject: [PATCH 0827/1291] feat(algebra/order/sub/defs): make `has_ordered_sub` a Prop (#18810) `has_ordered_sub` was incorrectly a Type, because of a Lean 3 bug. This is a backport of mathlib4 [#3436](https://github.com/leanprover-community/mathlib4/pull/3436) . Co-authored-by: Alex J Best --- src/algebra/order/sub/defs.lean | 2 +- src/data/real/nnreal.lean | 2 +- 2 files changed, 2 insertions(+), 2 deletions(-) diff --git a/src/algebra/order/sub/defs.lean b/src/algebra/order/sub/defs.lean index 53ad887104d00..e65e0290ed919 100644 --- a/src/algebra/order/sub/defs.lean +++ b/src/algebra/order/sub/defs.lean @@ -52,7 +52,7 @@ In other words, `a - b` is the least `c` such that `a ≤ b + c`. This is satisfied both by the subtraction in additive ordered groups and by truncated subtraction in canonically ordered monoids on many specific types. -/ -class has_ordered_sub (α : Type*) [has_le α] [has_add α] [has_sub α] := +class has_ordered_sub (α : Type*) [has_le α] [has_add α] [has_sub α] : Prop := (tsub_le_iff_right : ∀ a b c : α, a - b ≤ c ↔ a ≤ c + b) section has_add diff --git a/src/data/real/nnreal.lean b/src/data/real/nnreal.lean index bc8e5bf487e36..cca419cf48609 100644 --- a/src/data/real/nnreal.lean +++ b/src/data/real/nnreal.lean @@ -568,7 +568,7 @@ lemma sub_def {r p : ℝ≥0} : r - p = real.to_nnreal (r - p) := rfl lemma coe_sub_def {r p : ℝ≥0} : ↑(r - p) = max (r - p : ℝ) 0 := rfl -noncomputable example : has_ordered_sub ℝ≥0 := by apply_instance +example : has_ordered_sub ℝ≥0 := by apply_instance lemma sub_div (a b c : ℝ≥0) : (a - b) / c = a / c - b / c := tsub_div _ _ _ From e95e4f92c8f8da3c7f693c3ec948bcf9b6683f51 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 14 Apr 2023 22:49:11 +0000 Subject: [PATCH 0828/1291] feat(linear_algebra/finite_dimensional): generalize results to `module.finite` (#18811) This generalize the following from `finite_dimensional` over division rings to `module.finite` over free modules: * `finite_dimensional.nonempty_linear_equiv_of_finrank_eq` (moved from `nonempty_linear_equiv_of_finrank_eq`) * `finite_dimensional.nonempty_linear_equiv_iff_finrank_eq` (moved from `nonempty_linear_equiv_iff_finrank_eq`) * `linear_equiv.of_finrank_eq` * `module.finite.map` (moved from `finite_dimensional.submodule.map.finite_dimensional`). This is the only lemma moved across the porting tide. * `submodule.finrank_map_le` (moved from `finite_dimensional.finrank_map_le`) * `submodule.finrank_map_subtype_eq` (moved from `finite_dimensional.finrank_map_subtype_eq`, needs no finite or free assumptions at all) * `submodule.finrank_le_finrank_of_le` --- src/linear_algebra/bilinear_form.lean | 2 +- src/linear_algebra/finite_dimensional.lean | 51 +------------------ .../free_module/finite/rank.lean | 46 ++++++++++++++++- src/ring_theory/finiteness.lean | 4 ++ 4 files changed, 51 insertions(+), 52 deletions(-) diff --git a/src/linear_algebra/bilinear_form.lean b/src/linear_algebra/bilinear_form.lean index 2f3aeab57068b..c87fd45ebb10e 100644 --- a/src/linear_algebra/bilinear_form.lean +++ b/src/linear_algebra/bilinear_form.lean @@ -1181,7 +1181,7 @@ lemma finrank_add_finrank_orthogonal begin rw [← to_lin_restrict_ker_eq_inf_orthogonal _ _ b₁, ← to_lin_restrict_range_dual_coannihilator_eq_orthogonal _ _, - finrank_map_subtype_eq], + submodule.finrank_map_subtype_eq], conv_rhs { rw [← @subspace.finrank_add_finrank_dual_coannihilator_eq K V _ _ _ _ (B.to_lin.dom_restrict W).range, add_comm, ← add_assoc, add_comm (finrank K ↥((B.to_lin.dom_restrict W).ker)), diff --git a/src/linear_algebra/finite_dimensional.lean b/src/linear_algebra/finite_dimensional.lean index b72a40b849e35..efbac63192c3d 100644 --- a/src/linear_algebra/finite_dimensional.lean +++ b/src/linear_algebra/finite_dimensional.lean @@ -358,18 +358,7 @@ span_of_finite K $ s.finite_to_set /-- Pushforwards of finite-dimensional submodules are finite-dimensional. -/ instance (f : V →ₗ[K] V₂) (p : submodule K V) [h : finite_dimensional K p] : finite_dimensional K (p.map f) := -begin - unfreezingI { rw [finite_dimensional, ← iff_fg, is_noetherian.iff_rank_lt_aleph_0] at h ⊢ }, - rw [← cardinal.lift_lt.{v' v}], - rw [← cardinal.lift_lt.{v v'}] at h, - rw [cardinal.lift_aleph_0] at h ⊢, - exact (lift_rank_map_le f p).trans_lt h -end - -/-- Pushforwards of finite-dimensional submodules have a smaller finrank. -/ -lemma finrank_map_le (f : V →ₗ[K] V₂) (p : submodule K V) [finite_dimensional K p] : - finrank K (p.map f) ≤ finrank K p := -by simpa [← finrank_eq_rank', -finrank_eq_rank] using lift_rank_map_le f p +module.finite.map _ _ variable {K} @@ -781,32 +770,6 @@ 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. --/ -theorem nonempty_linear_equiv_of_finrank_eq [finite_dimensional K V] [finite_dimensional K V₂] - (cond : finrank K V = finrank K V₂) : nonempty (V ≃ₗ[K] V₂) := -nonempty_linear_equiv_of_lift_rank_eq $ by simp only [← finrank_eq_rank, cond, lift_nat_cast] - -/-- -Two finite-dimensional vector spaces are isomorphic if and only if they have the same (finite) -dimension. --/ -theorem nonempty_linear_equiv_iff_finrank_eq [finite_dimensional K V] [finite_dimensional K V₂] : - nonempty (V ≃ₗ[K] V₂) ↔ finrank K V = finrank K V₂ := -⟨λ ⟨h⟩, h.finrank_eq, λ h, nonempty_linear_equiv_of_finrank_eq h⟩ - -variables (V V₂) - -/-- -Two finite-dimensional vector spaces are isomorphic if they have the same (finite) dimension. --/ -noncomputable def linear_equiv.of_finrank_eq [finite_dimensional K V] [finite_dimensional K V₂] - (cond : finrank K V = finrank K V₂) : V ≃ₗ[K] V₂ := -classical.choice $ nonempty_linear_equiv_of_finrank_eq cond - -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₂ := begin @@ -821,12 +784,7 @@ lemma eq_of_le_of_finrank_eq {S₁ S₂ : submodule K V} [finite_dimensional K S (hd : finrank K S₁ = finrank K S₂) : S₁ = S₂ := eq_of_le_of_finrank_le hle hd.ge -@[simp] -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 - -variables {V₂} [finite_dimensional K V] [finite_dimensional K V₂] +variables [finite_dimensional K V] [finite_dimensional K V₂] /-- Given isomorphic subspaces `p q` of vector spaces `V` and `V₁` respectively, `p.quotient` is isomorphic to `q.quotient`. -/ @@ -1054,11 +1012,6 @@ lemma eq_top_of_finrank_eq [finite_dimensional K V] {S : submodule K V} (h : finrank K S = finrank K V) : S = ⊤ := finite_dimensional.eq_of_le_of_finrank_eq le_top (by simp [h, finrank_top]) -lemma finrank_le_finrank_of_le {s t : submodule K V} [finite_dimensional K t] - (hst : s ≤ t) : finrank K s ≤ finrank K t := -calc finrank K s = finrank K (comap t.subtype s) : (comap_subtype_equiv_of_le hst).finrank_eq.symm -... ≤ finrank K t : finrank_le _ - lemma finrank_mono [finite_dimensional K V] : monotone (λ (s : submodule K V), finrank K s) := λ s t, finrank_le_finrank_of_le diff --git a/src/linear_algebra/free_module/finite/rank.lean b/src/linear_algebra/free_module/finite/rank.lean index bc0b31beaddc5..1b2ace77c9a8b 100644 --- a/src/linear_algebra/free_module/finite/rank.lean +++ b/src/linear_algebra/free_module/finite/rank.lean @@ -30,6 +30,18 @@ namespace finite_dimensional open module.free section ring +variables [ring R] +variables [add_comm_group M] [module R M] +variables [add_comm_group N] [module R N] + +@[simp] +lemma submodule.finrank_map_subtype_eq (p : submodule R M) (q : submodule R p) : + finrank R (q.map p.subtype) = finrank R q := +(submodule.equiv_subtype_map p q).symm.finrank_eq + +end ring + +section ring_finite variables [ring R] [strong_rank_condition R] variables [add_comm_group M] [module R M] [module.finite R M] @@ -49,7 +61,7 @@ end @[simp] lemma finrank_eq_rank : ↑(finrank R M) = module.rank R M := by { rw [finrank, cast_to_nat_of_lt_aleph_0 (rank_lt_aleph_0 R M)] } -end ring +end ring_finite section ring_free @@ -104,6 +116,25 @@ lemma finrank_matrix (m n : Type*) [fintype m] [fintype n] : finrank R (matrix m n R) = (card m) * (card n) := by { simp [finrank] } +variables {R M N} + +/-- Two finite and free modules are isomorphic if they have the same (finite) rank. -/ +theorem nonempty_linear_equiv_of_finrank_eq + (cond : finrank R M = finrank R N) : nonempty (M ≃ₗ[R] N) := +nonempty_linear_equiv_of_lift_rank_eq $ by simp only [← finrank_eq_rank, cond, lift_nat_cast] + +/-- Two finite and free modules are isomorphic if and only if they have the same (finite) rank. -/ +theorem nonempty_linear_equiv_iff_finrank_eq : + nonempty (M ≃ₗ[R] N) ↔ finrank R M = finrank R N := +⟨λ ⟨h⟩, h.finrank_eq, λ h, nonempty_linear_equiv_of_finrank_eq h⟩ + +variables (M N) + +/-- Two finite and free modules are isomorphic if they have the same (finite) rank. -/ +noncomputable def _root_.linear_equiv.of_finrank_eq (cond : finrank R M = finrank R N) : + M ≃ₗ[R] N := +classical.choice $ nonempty_linear_equiv_of_finrank_eq cond + end ring_free section comm_ring @@ -115,7 +146,7 @@ variables [add_comm_group N] [module R N] [module.free R N] [module.finite R N] /-- The finrank of `M ⊗[R] N` is `(finrank R M) * (finrank R N)`. -/ @[simp] lemma finrank_tensor_product (M : Type v) (N : Type w) [add_comm_group M] [module R M] [module.free R M] [add_comm_group N] [module R N] [module.free R N] : -finrank R (M ⊗[R] N) = (finrank R M) * (finrank R N) := + finrank R (M ⊗[R] N) = (finrank R M) * (finrank R N) := by { simp [finrank] } end comm_ring @@ -151,4 +182,15 @@ lemma submodule.finrank_quotient_le [module.finite R M] (s : submodule R M) : by simpa only [cardinal.to_nat_lift] using to_nat_le_of_le_of_lt_aleph_0 (rank_lt_aleph_0 _ _) ((submodule.mkq s).rank_le_of_surjective (surjective_quot_mk _)) +/-- Pushforwards of finite submodules have a smaller finrank. -/ +lemma submodule.finrank_map_le (f : M →ₗ[R] N) (p : submodule R M) [module.finite R p] : + finrank R (p.map f) ≤ finrank R p := +finrank_le_finrank_of_rank_le_rank (lift_rank_map_le _ _) (rank_lt_aleph_0 _ _) + +lemma submodule.finrank_le_finrank_of_le {s t : submodule R M} [module.finite R t] + (hst : s ≤ t) : finrank R s ≤ finrank R t := +calc finrank R s = finrank R (s.comap t.subtype) + : (submodule.comap_subtype_equiv_of_le hst).finrank_eq.symm +... ≤ finrank R t : submodule.finrank_le _ + end diff --git a/src/ring_theory/finiteness.lean b/src/ring_theory/finiteness.lean index 4f3d102acda03..1b22d2e24e4f0 100644 --- a/src/ring_theory/finiteness.lean +++ b/src/ring_theory/finiteness.lean @@ -479,6 +479,10 @@ end⟩ instance range [finite R M] (f : M →ₗ[R] N) : finite R f.range := of_surjective f.range_restrict $ λ ⟨x, y, hy⟩, ⟨y, subtype.ext hy⟩ +/-- Pushforwards of finite submodules are finite. -/ +instance map (p : submodule R M) [finite R p] (f : M →ₗ[R] N) : finite R (p.map f) := +of_surjective (f.restrict $ λ _, mem_map_of_mem) $ λ ⟨x, y, hy, hy'⟩, ⟨⟨_, hy⟩, subtype.ext hy'⟩ + variables (R) instance self : finite R R := From 4f4a1c875d0baa92ab5d92f3fb1bb258ad9f3e5b Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Sat, 15 Apr 2023 04:32:04 +0000 Subject: [PATCH 0829/1291] chore(*): add mathlib4 synchronization comments (#18807) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.category.Module.epi_mono` * `algebraic_topology.split_simplicial_object` * `analysis.asymptotics.specific_asymptotics` * `analysis.asymptotics.superpolynomial_decay` * `analysis.asymptotics.theta` * `analysis.box_integral.box.basic` * `analysis.locally_convex.balanced_core_hull` * `analysis.normed_space.extr` * `category_theory.limits.over` * `category_theory.sites.sheafification` * `linear_algebra.free_module.rank` * `order.filter.zero_and_bounded_at_filter` * `topology.category.Top.epi_mono` --- src/algebra/category/Module/epi_mono.lean | 3 +++ src/algebraic_topology/split_simplicial_object.lean | 3 +++ src/analysis/asymptotics/specific_asymptotics.lean | 3 +++ src/analysis/asymptotics/superpolynomial_decay.lean | 3 +++ src/analysis/asymptotics/theta.lean | 3 +++ src/analysis/box_integral/box/basic.lean | 3 +++ src/analysis/locally_convex/balanced_core_hull.lean | 3 +++ src/analysis/normed_space/extr.lean | 3 +++ src/category_theory/limits/over.lean | 3 +++ src/category_theory/sites/sheafification.lean | 3 +++ src/linear_algebra/free_module/rank.lean | 3 +++ src/order/filter/zero_and_bounded_at_filter.lean | 3 +++ src/topology/category/Top/epi_mono.lean | 3 +++ 13 files changed, 39 insertions(+) diff --git a/src/algebra/category/Module/epi_mono.lean b/src/algebra/category/Module/epi_mono.lean index fb24e19c6092d..88a5c59967586 100644 --- a/src/algebra/category/Module/epi_mono.lean +++ b/src/algebra/category/Module/epi_mono.lean @@ -9,6 +9,9 @@ import algebra.category.Module.basic /-! # Monomorphisms in `Module R` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file shows that an `R`-linear map is a monomorphism in the category of `R`-modules if and only if it is injective, and similarly an epimorphism if and only if it is surjective. -/ diff --git a/src/algebraic_topology/split_simplicial_object.lean b/src/algebraic_topology/split_simplicial_object.lean index d82c9fbb11d2d..18a1a01d2a18c 100644 --- a/src/algebraic_topology/split_simplicial_object.lean +++ b/src/algebraic_topology/split_simplicial_object.lean @@ -11,6 +11,9 @@ import category_theory.limits.shapes.finite_products # Split simplicial objects +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we introduce the notion of split simplicial object. If `C` is a category that has finite coproducts, a splitting `s : splitting X` of a simplical object `X` in `C` consists diff --git a/src/analysis/asymptotics/specific_asymptotics.lean b/src/analysis/asymptotics/specific_asymptotics.lean index 33909a5bc401e..b98632ff3194e 100644 --- a/src/analysis/asymptotics/specific_asymptotics.lean +++ b/src/analysis/asymptotics/specific_asymptotics.lean @@ -9,6 +9,9 @@ import analysis.asymptotics.asymptotics /-! # A collection of specific asymptotic results +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains specific lemmas about asymptotics which don't have their place in the general theory developped in `analysis.asymptotics.asymptotics`. -/ diff --git a/src/analysis/asymptotics/superpolynomial_decay.lean b/src/analysis/asymptotics/superpolynomial_decay.lean index bf461b0a82afe..3f003aa64537c 100644 --- a/src/analysis/asymptotics/superpolynomial_decay.lean +++ b/src/analysis/asymptotics/superpolynomial_decay.lean @@ -11,6 +11,9 @@ import topology.algebra.order.liminf_limsup /-! # Super-Polynomial Function Decay +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines a predicate `asymptotics.superpolynomial_decay f` for a function satisfying one of following equivalent definitions (The definition is in terms of the first condition): diff --git a/src/analysis/asymptotics/theta.lean b/src/analysis/asymptotics/theta.lean index fb18642e27344..98375933808fc 100644 --- a/src/analysis/asymptotics/theta.lean +++ b/src/analysis/asymptotics/theta.lean @@ -8,6 +8,9 @@ import analysis.asymptotics.asymptotics /-! # Asymptotic equivalence up to a constant +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define `asymptotics.is_Theta l f g` (notation: `f =Θ[l] g`) as `f =O[l] g ∧ g =O[l] f`, then prove basic properties of this equivalence relation. -/ diff --git a/src/analysis/box_integral/box/basic.lean b/src/analysis/box_integral/box/basic.lean index e9c525883e371..71807a3aa84a4 100644 --- a/src/analysis/box_integral/box/basic.lean +++ b/src/analysis/box_integral/box/basic.lean @@ -10,6 +10,9 @@ import topology.metric_space.basic /-! # Rectangular boxes in `ℝⁿ` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define rectangular boxes in `ℝⁿ`. As usual, we represent `ℝⁿ` as the type of functions `ι → ℝ` (usually `ι = fin n` for some `n`). When we need to interpret a box `[l, u]` as a set, we use the product `{x | ∀ i, l i < x i ∧ x i ≤ u i}` of half-open intervals `(l i, u i]`. We diff --git a/src/analysis/locally_convex/balanced_core_hull.lean b/src/analysis/locally_convex/balanced_core_hull.lean index b8d59634f79aa..0533613a15d02 100644 --- a/src/analysis/locally_convex/balanced_core_hull.lean +++ b/src/analysis/locally_convex/balanced_core_hull.lean @@ -8,6 +8,9 @@ import analysis.locally_convex.basic /-! # Balanced Core and Balanced Hull +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions * `balanced_core`: The largest balanced subset of a set `s`. diff --git a/src/analysis/normed_space/extr.lean b/src/analysis/normed_space/extr.lean index b458f8d5dc8aa..29b2bc70da3a9 100644 --- a/src/analysis/normed_space/extr.lean +++ b/src/analysis/normed_space/extr.lean @@ -9,6 +9,9 @@ import topology.local_extr /-! # (Local) maximums in a normed space +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove the following lemma, see `is_max_filter.norm_add_same_ray`. If `f : α → E` is a function such that `norm ∘ f` has a maximum along a filter `l` at a point `c` and `y` is a vector on the same ray as `f c`, then the function `λ x, ‖f x + y‖` has a maximul along `l` at `c`. diff --git a/src/category_theory/limits/over.lean b/src/category_theory/limits/over.lean index 16be8d564aece..ca2272a529949 100644 --- a/src/category_theory/limits/over.lean +++ b/src/category_theory/limits/over.lean @@ -13,6 +13,9 @@ import category_theory.limits.comma /-! # Limits and colimits in the over and under categories +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Show that the forgetful functor `forget X : over X ⥤ C` creates colimits, and hence `over X` has any colimits that `C` has (as well as the dual that `forget X : under X ⟶ C` creates limits). diff --git a/src/category_theory/sites/sheafification.lean b/src/category_theory/sites/sheafification.lean index 557577a7291c9..f424ae433bc56 100644 --- a/src/category_theory/sites/sheafification.lean +++ b/src/category_theory/sites/sheafification.lean @@ -12,6 +12,9 @@ import category_theory.concrete_category.elementwise # Sheafification +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We construct the sheafification of a presheaf over a site `C` with values in `D` whenever `D` is a concrete category for which the forgetful functor preserves the appropriate (co)limits and reflects isomorphisms. diff --git a/src/linear_algebra/free_module/rank.lean b/src/linear_algebra/free_module/rank.lean index 875e52ed714a7..0910b9caf26f9 100644 --- a/src/linear_algebra/free_module/rank.lean +++ b/src/linear_algebra/free_module/rank.lean @@ -10,6 +10,9 @@ import linear_algebra.dimension # Extra results about `module.rank` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains some extra results not in `linear_algebra.dimension`. -/ diff --git a/src/order/filter/zero_and_bounded_at_filter.lean b/src/order/filter/zero_and_bounded_at_filter.lean index ef69104312417..5cc4433ee6424 100644 --- a/src/order/filter/zero_and_bounded_at_filter.lean +++ b/src/order/filter/zero_and_bounded_at_filter.lean @@ -10,6 +10,9 @@ import analysis.asymptotics.asymptotics /-! # Zero and Bounded at filter +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given a filter `l` we define the notion of a function being `zero_at_filter` as well as being `bounded_at_filter`. Alongside this we construct the `submodule`, `add_submonoid` of functions that are `zero_at_filter`. Similarly, we construct the `submodule` and `subalgebra` of functions diff --git a/src/topology/category/Top/epi_mono.lean b/src/topology/category/Top/epi_mono.lean index fd7100d1ebe59..2ac9b10b9a059 100644 --- a/src/topology/category/Top/epi_mono.lean +++ b/src/topology/category/Top/epi_mono.lean @@ -8,6 +8,9 @@ import topology.category.Top.adjunctions /-! # Epi- and monomorphisms in `Top` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file shows that a continuous function is an epimorphism in the category of topological spaces if and only if it is surjective, and that a continuous function is a monomorphism in the category of topological spaces if and only if it is injective. From 17ad94b4953419f3e3ce3e77da3239c62d1d09f0 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sun, 16 Apr 2023 20:29:56 +0000 Subject: [PATCH 0830/1291] ci: fix bibtool installation (#18816) Adding `sudo apt-get update` seems to fix the installation failure; perhaps our package list referred to an old version that the default mirror decided not to bother hosting any more. Co-authored-by: MohanadAhmed --- .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, 12 insertions(+), 4 deletions(-) diff --git a/.github/workflows/bors.yml b/.github/workflows/bors.yml index 8b13cb25e519b..0b8033c11e328 100644 --- a/.github/workflows/bors.yml +++ b/.github/workflows/bors.yml @@ -36,7 +36,9 @@ jobs: - name: Install bibtool if: ${{ 'bors' == 'ubuntu-latest' }} - run: sudo apt install -y bibtool + run: | + sudo apt-get update + sudo apt-get install -y bibtool - name: install Python if: ${{ 'bors' == 'ubuntu-latest' }} diff --git a/.github/workflows/build.yml b/.github/workflows/build.yml index ad8edebc51b94..b61b8eea978f3 100644 --- a/.github/workflows/build.yml +++ b/.github/workflows/build.yml @@ -44,7 +44,9 @@ jobs: - name: Install bibtool if: ${{ 'ubuntu-latest' == 'ubuntu-latest' }} - run: sudo apt install -y bibtool + run: | + sudo apt-get update + sudo apt-get install -y bibtool - name: install Python if: ${{ 'ubuntu-latest' == 'ubuntu-latest' }} diff --git a/.github/workflows/build.yml.in b/.github/workflows/build.yml.in index b9ebf269ddafe..a3f6edde6168c 100644 --- a/.github/workflows/build.yml.in +++ b/.github/workflows/build.yml.in @@ -22,7 +22,9 @@ jobs: - name: Install bibtool if: ${{ 'STYLE_LINT_RUNNER' == 'ubuntu-latest' }} - run: sudo apt install -y bibtool + run: | + sudo apt-get update + sudo apt-get install -y bibtool - name: install Python if: ${{ 'STYLE_LINT_RUNNER' == 'ubuntu-latest' }} diff --git a/.github/workflows/build_fork.yml b/.github/workflows/build_fork.yml index 58d0322673660..91ad14c848268 100644 --- a/.github/workflows/build_fork.yml +++ b/.github/workflows/build_fork.yml @@ -42,7 +42,9 @@ jobs: - name: Install bibtool if: ${{ 'ubuntu-latest' == 'ubuntu-latest' }} - run: sudo apt install -y bibtool + run: | + sudo apt-get update + sudo apt-get install -y bibtool - name: install Python if: ${{ 'ubuntu-latest' == 'ubuntu-latest' }} From 95e83ced9542828815f53a1096a4d373c1b08a77 Mon Sep 17 00:00:00 2001 From: Kevin Buzzard Date: Mon, 17 Apr 2023 00:33:14 +0000 Subject: [PATCH 0831/1291] feat(category_theory/site/limits): generalise universes (#18817) The objects in category `C` in this file had type `Type max v u` but there's no reason they can't just have type `Type u`. There are no uses in this file of `u` as a universe by itself, so this does not make any results less general. This makes porting the file easier. [Zulip discussion](https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/Universe.20issues.20with.20concrete.20categories/near/350340297) --- src/category_theory/sites/limits.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/category_theory/sites/limits.lean b/src/category_theory/sites/limits.lean index 50ff8db9774df..49efac80318e1 100644 --- a/src/category_theory/sites/limits.lean +++ b/src/category_theory/sites/limits.lean @@ -34,7 +34,7 @@ open opposite section limits universes w v u z -variables {C : Type (max v u)} [category.{v} C] {J : grothendieck_topology C} +variables {C : Type u} [category.{v} C] {J : grothendieck_topology C} variables {D : Type w} [category.{max v u} D] variables {K : Type z} [small_category K] @@ -166,7 +166,7 @@ end limits section colimits universes w v u -variables {C : Type (max v u)} [category.{v} C] {J : grothendieck_topology C} +variables {C : Type u} [category.{v} C] {J : grothendieck_topology C} variables {D : Type w} [category.{max v u} D] variables {K : Type (max v u)} [small_category K] -- Now we need a handful of instances to obtain sheafification... From a968611b6a772cf7bdf61146e6d62fc882c92372 Mon Sep 17 00:00:00 2001 From: Moritz Doll Date: Mon, 17 Apr 2023 06:47:49 +0000 Subject: [PATCH 0832/1291] feat(analysis/schwartz_space): lemmas for supremum of seminorms (#18648) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The main result is the bound `one_add_le_seminorm_sup_apply`, which is sometimes in the literature used as the definition of the Schwartz space. We don't really care about the `2^k` factor, since this result is usually used to prove that certain operators are bounded on Schwartz space and hence finiteness of the right hand side is all that matters. One application is to show that the product of a Schwartz function and a smooth polynomially growing function is again Schwartz. Co-authored-by: Yaël Dillies --- src/analysis/schwartz_space.lean | 51 ++++++++++++++++++++++++++------ src/data/finset/lattice.lean | 8 +++++ 2 files changed, 50 insertions(+), 9 deletions(-) diff --git a/src/analysis/schwartz_space.lean b/src/analysis/schwartz_space.lean index f9ac9cea0cd27..ccfc4a366a59e 100644 --- a/src/analysis/schwartz_space.lean +++ b/src/analysis/schwartz_space.lean @@ -23,7 +23,7 @@ smooth functions `f : E → F`, where `E` and `F` are real normed vector spaces natural numbers `k` and `n` we have uniform bounds `‖x‖^k * ‖iterated_fderiv ℝ n f x‖ < C`. This approach completely avoids using partial derivatives as well as polynomials. We construct the topology on the Schwartz space by a family of seminorms, which are the best -constants in the above estimates, which is by abstract theory from +constants in the above estimates. The abstract theory of topological vector spaces developed in `seminorm_family.module_filter_basis` and `with_seminorms.to_locally_convex_space` turns the Schwartz space into a locally convex topological vector space. @@ -41,6 +41,8 @@ decay faster than any power of `‖x‖`. * `schwartz_map.uniform_add_group` and `schwartz_map.locally_convex`: The Schwartz space is a locally convex topological vector space. +* `schwartz_map.one_add_le_sup_seminorm_apply`: For a Schwartz function `f` there is a uniform bound +on `(1 + ‖x‖) ^ k * ‖iterated_fderiv ℝ n f x‖`. ## Implementation details @@ -430,18 +432,11 @@ begin rwa [pow_zero, one_mul] at this, end -end seminorms - -section topology - -/-! ### The topology on the Schwartz space-/ - -variables [normed_field 𝕜] [normed_space 𝕜 F] [smul_comm_class ℝ 𝕜 F] variables (𝕜 E F) /-- The family of Schwartz seminorms. -/ def _root_.schwartz_seminorm_family : seminorm_family 𝕜 𝓢(E, F) (ℕ × ℕ) := -λ n, seminorm 𝕜 n.1 n.2 +λ m, seminorm 𝕜 m.1 m.2 @[simp] lemma schwartz_seminorm_family_apply (n k : ℕ) : schwartz_seminorm_family 𝕜 E F (n,k) = schwartz_map.seminorm 𝕜 n k := rfl @@ -449,6 +444,44 @@ def _root_.schwartz_seminorm_family : seminorm_family 𝕜 𝓢(E, F) (ℕ × @[simp] lemma schwartz_seminorm_family_apply_zero : schwartz_seminorm_family 𝕜 E F 0 = schwartz_map.seminorm 𝕜 0 0 := rfl +variables {𝕜 E F} + +/-- A more convenient version of `le_sup_seminorm_apply`. + +The set `finset.Iic m` is the set of all pairs `(k', n')` with `k' ≤ m.1` and `n' ≤ m.2`. +Note that the constant is far from optimal. -/ +lemma one_add_le_sup_seminorm_apply {m : ℕ × ℕ} {k n : ℕ} (hk : k ≤ m.1) (hn : n ≤ m.2) + (f : 𝓢(E, F)) (x : E) : + (1 + ‖x‖) ^ k * ‖iterated_fderiv ℝ n f x‖ + ≤ 2^m.1 * (finset.Iic m).sup (λ m, seminorm 𝕜 m.1 m.2) f := +begin + rw [add_comm, add_pow], + simp only [one_pow, mul_one, finset.sum_congr, finset.sum_mul], + norm_cast, + rw ← nat.sum_range_choose m.1, + push_cast, + rw [finset.sum_mul], + have hk' : finset.range (k + 1) ⊆ finset.range (m.1 + 1) := + by rwa [finset.range_subset, add_le_add_iff_right], + refine le_trans (finset.sum_le_sum_of_subset_of_nonneg hk' (λ _ _ _, by positivity)) _, + refine finset.sum_le_sum (λ i hi, _), + rw [mul_comm (‖x‖^i), mul_assoc], + refine mul_le_mul _ _ (by positivity) (by positivity), + { norm_cast, + exact i.choose_le_choose hk }, + exact (le_seminorm 𝕜 i n f x).trans (seminorm.le_def.1 (finset.le_sup_of_le + (finset.mem_Iic.2 $ prod.mk_le_mk.2 ⟨finset.mem_range_succ_iff.mp hi, hn⟩) le_rfl) _), +end + +end seminorms + +section topology + +/-! ### The topology on the Schwartz space-/ + +variables [normed_field 𝕜] [normed_space 𝕜 F] [smul_comm_class ℝ 𝕜 F] +variables (𝕜 E F) + instance : topological_space 𝓢(E, F) := (schwartz_seminorm_family ℝ E F).module_filter_basis.topology' diff --git a/src/data/finset/lattice.lean b/src/data/finset/lattice.lean index 7b991d9957df4..c44107febc9ca 100644 --- a/src/data/finset/lattice.lean +++ b/src/data/finset/lattice.lean @@ -83,6 +83,8 @@ lemma sup_const_le : s.sup (λ _, a) ≤ a := finset.sup_le $ λ _ _, le_rfl lemma le_sup {b : β} (hb : b ∈ s) : f b ≤ s.sup f := finset.sup_le_iff.1 le_rfl _ hb +lemma le_sup_of_le {b : β} (hb : b ∈ s) (h : a ≤ f b) : a ≤ s.sup f := h.trans $ le_sup hb + @[simp] lemma sup_bUnion [decidable_eq β] (s : finset γ) (t : γ → finset β) : (s.bUnion t).sup f = s.sup (λ x, (t x).sup f) := eq_of_forall_ge_iff $ λ c, by simp [@forall_swap _ β] @@ -310,6 +312,8 @@ lemma le_inf_const_le : a ≤ s.inf (λ _, a) := finset.le_inf $ λ _ _, le_rfl lemma inf_le {b : β} (hb : b ∈ s) : s.inf f ≤ f b := finset.le_inf_iff.1 le_rfl _ hb +lemma inf_le_of_le {b : β} (hb : b ∈ s) (h : f b ≤ a) : s.inf f ≤ a := (inf_le hb).trans h + lemma inf_mono_fun {g : β → α} (h : ∀ b ∈ s, f b ≤ g b) : s.inf f ≤ s.inf g := finset.le_inf (λ b hb, le_trans (inf_le hb) (h b hb)) @@ -542,6 +546,9 @@ by { rw [←with_bot.coe_le_coe, coe_sup'], lemma le_sup' {b : β} (h : b ∈ s) : f b ≤ s.sup' ⟨b, h⟩ f := by { rw [←with_bot.coe_le_coe, coe_sup'], exact le_sup h, } +lemma le_sup'_of_le {a : α} {b : β} (hb : b ∈ s) (h : a ≤ f b) : a ≤ s.sup' ⟨b, hb⟩ f := +h.trans $ le_sup' _ hb + @[simp] lemma sup'_const (a : α) : s.sup' H (λ b, a) = a := begin apply le_antisymm, @@ -637,6 +644,7 @@ variables {s : finset β} (H : s.nonempty) (f : β → α) {a : α} {b : β} 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 +lemma inf'_le_of_le (hb : b ∈ s) (h : f b ≤ a) : s.inf' ⟨b, hb⟩ f ≤ a := (inf'_le _ hb).trans h @[simp] lemma inf'_const (a : α) : s.inf' H (λ b, a) = a := @sup'_const αᵒᵈ _ _ _ H _ From 991ff3b5269848f6dd942ae8e9dd3c946035dc8b Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Mon, 17 Apr 2023 10:35:23 +0000 Subject: [PATCH 0833/1291] golf(set_theory/ordinal/cantor_normal_form): golf theorems (#16009) We move `div_opow_log_pos` out of a proof and open the `list` namespace. Mathlib 4: https://github.com/leanprover-community/mathlib4/pull/3189 --- .../ordinal/cantor_normal_form.lean | 24 +++++++------------ 1 file changed, 8 insertions(+), 16 deletions(-) diff --git a/src/set_theory/ordinal/cantor_normal_form.lean b/src/set_theory/ordinal/cantor_normal_form.lean index b084cd7f25f8c..987fb56b66212 100644 --- a/src/set_theory/ordinal/cantor_normal_form.lean +++ b/src/set_theory/ordinal/cantor_normal_form.lean @@ -36,7 +36,7 @@ noncomputable theory universe u -open order +open list namespace ordinal @@ -90,16 +90,15 @@ by simp [CNF_ne_zero ho, log_eq_zero hb] /-- Evaluating the Cantor normal form of an ordinal returns the ordinal. -/ theorem CNF_foldr (b o : ordinal) : (CNF b o).foldr (λ p r, b ^ p.1 * p.2 + r) 0 = o := CNF_rec b (by { rw CNF_zero, refl }) - (λ o ho IH, by rw [CNF_ne_zero ho, list.foldr_cons, IH, div_add_mod]) o + (λ o ho IH, by rw [CNF_ne_zero ho, foldr_cons, IH, div_add_mod]) o /-- Every exponent in the Cantor normal form `CNF b o` is less or equal to `log b o`. -/ theorem CNF_fst_le_log {b o : ordinal.{u}} {x : ordinal × ordinal} : x ∈ CNF b o → x.1 ≤ log b o := begin refine CNF_rec b _ (λ o ho H, _) o, - { rw CNF_zero, - exact false.elim }, - { rw [CNF_ne_zero ho, list.mem_cons_iff], + { simp }, + { rw [CNF_ne_zero ho, mem_cons_iff], rintro (rfl | h), { exact le_rfl }, { exact (H h).trans (log_mono_right _ (mod_opow_log_lt_self b ho).le) } } @@ -114,17 +113,10 @@ theorem CNF_lt_snd {b o : ordinal.{u}} {x : ordinal × ordinal} : x ∈ CNF b o begin refine CNF_rec b _ (λ o ho IH, _) o, { simp }, - { rcases eq_zero_or_pos b with rfl | hb, - { rw [zero_CNF ho, list.mem_singleton], - rintro rfl, - exact ordinal.pos_iff_ne_zero.2 ho }, - { rw CNF_ne_zero ho, - rintro (rfl | h), - { simp, - rw div_pos, - { exact opow_log_le_self _ ho }, - { exact (opow_pos _ hb).ne' } }, - { exact IH h } } } + { rw CNF_ne_zero ho, + rintro (rfl | h), + { exact div_opow_log_pos b ho }, + { exact IH h } } end /-- Every coefficient in the Cantor normal form `CNF b o` is less than `b`. -/ From 155d5519569cecf21f48c534da8b729890e20ce6 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Mon, 17 Apr 2023 11:40:52 +0000 Subject: [PATCH 0834/1291] feat(algebra/module/submodule/basic): add has_vadd (#18815) --- src/algebra/module/submodule/basic.lean | 32 +++++++++++++++++++++++++ 1 file changed, 32 insertions(+) diff --git a/src/algebra/module/submodule/basic.lean b/src/algebra/module/submodule/basic.lean index 7e9547166f1ad..39aa53575d85d 100644 --- a/src/algebra/module/submodule/basic.lean +++ b/src/algebra/module/submodule/basic.lean @@ -257,6 +257,38 @@ lemma injective_subtype : injective p.subtype := subtype.coe_injective @[simp] lemma coe_sum (x : ι → p) (s : finset ι) : ↑(∑ i in s, x i) = ∑ i in s, (x i : M) := map_sum p.subtype _ _ +section add_action + +/-! ### Additive actions by `submodule`s + +These instances transfer the action by an element `m : M` of a `R`-module `M` written as `m +ᵥ a` +onto the action by an element `s : S` of a submodule `S : submodule R M` such that +`s +ᵥ a = (s : M) +ᵥ a`. + +These instances work particularly well in conjunction with `add_group.to_add_action`, enabling +`s +ᵥ m` as an alias for `↑s + m`. + +-/ + +variables {α β : Type*} + +instance [has_vadd M α] : has_vadd p α := p.to_add_submonoid.has_vadd + +instance vadd_comm_class [has_vadd M β] [has_vadd α β] [vadd_comm_class M α β] : + vadd_comm_class p α β := ⟨λ a, (vadd_comm (a : M) : _)⟩ + +instance [has_vadd M α] [has_faithful_vadd M α] : + has_faithful_vadd p α := ⟨λ x y h, subtype.ext $ eq_of_vadd_eq_vadd h⟩ + +/-- The action by a submodule is the action by the underlying module. -/ +instance [add_action M α] : add_action p α := add_action.comp_hom _ p.subtype.to_add_monoid_hom + +variable {p} + +lemma vadd_def [has_vadd M α] (g : p) (m : α) : g +ᵥ m = (g : M) +ᵥ m := rfl + +end add_action + section restrict_scalars variables (S) [semiring S] [module S M] [module R M] [has_smul S R] [is_scalar_tower S R M] From 13bf7613c96a9fd66a81b9020a82cad9a6ea1fcf Mon Sep 17 00:00:00 2001 From: sgouezel Date: Mon, 17 Apr 2023 16:24:10 +0000 Subject: [PATCH 0835/1291] feat(measure_theory/functions/lp_space): bounds on the sum of functions in L^p, p < 1, and applications (#18796) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The `L^p` space satisfies the triangular inequality for `p ≥ 1`. We show that, for `p < 1`, it satisfies a weaker inequality (with the loss of a multiplicative constant `2^(1/p - 1)`), which is still enough for several results. This makes it possible to remove the assumptions on `p` in results on density of continuous functions. --- src/analysis/mean_inequalities_pow.lean | 36 ++++- .../function/continuous_map_dense.lean | 119 ++++++-------- src/measure_theory/function/lp_space.lean | 152 ++++++++++++------ .../function/simple_func_dense_lp.lean | 106 +++++++----- .../integral/mean_inequalities.lean | 23 +++ 5 files changed, 263 insertions(+), 173 deletions(-) diff --git a/src/analysis/mean_inequalities_pow.lean b/src/analysis/mean_inequalities_pow.lean index 57f964e83f40d..8292da30dfac0 100644 --- a/src/analysis/mean_inequalities_pow.lean +++ b/src/analysis/mean_inequalities_pow.lean @@ -109,6 +109,20 @@ begin { simp [hw', fin.sum_univ_succ], }, end +/-- Unweighted mean inequality, version for two elements of `ℝ≥0` and real exponents. -/ +theorem rpow_add_le_mul_rpow_add_rpow (z₁ z₂ : ℝ≥0) {p : ℝ} (hp : 1 ≤ p) : + (z₁ + z₂) ^ p ≤ 2^(p-1) * (z₁ ^ p + z₂ ^ p) := +begin + rcases eq_or_lt_of_le hp with rfl|h'p, + { simp only [rpow_one, sub_self, rpow_zero, one_mul] }, + convert rpow_arith_mean_le_arith_mean2_rpow (1/2) (1/2) (2 * z₁) (2 * z₂) (add_halves 1) hp, + { simp only [one_div, inv_mul_cancel_left₀, ne.def, bit0_eq_zero, one_ne_zero, not_false_iff] }, + { simp only [one_div, inv_mul_cancel_left₀, ne.def, bit0_eq_zero, one_ne_zero, not_false_iff] }, + { have A : p - 1 ≠ 0 := ne_of_gt (sub_pos.2 h'p), + simp only [mul_rpow, rpow_sub' _ A, div_eq_inv_mul, rpow_one, mul_one], + ring } +end + /-- Weighted generalized mean inequality, version for sums over finite sets, with `ℝ≥0`-valued functions and real exponents. -/ theorem arith_mean_le_rpow_mean (w z : ι → ℝ≥0) (hw' : ∑ i in s, w i = 1) {p : ℝ} @@ -117,10 +131,6 @@ theorem arith_mean_le_rpow_mean (w z : ι → ℝ≥0) (hw' : ∑ i in s, w i = by exact_mod_cast real.arith_mean_le_rpow_mean s _ _ (λ i _, (w i).coe_nonneg) (by exact_mod_cast hw') (λ i _, (z i).coe_nonneg) hp -end nnreal - -namespace nnreal - private lemma add_rpow_le_one_of_add_le_one {p : ℝ} (a b : ℝ≥0) (hab : a + b ≤ 1) (hp1 : 1 ≤ p) : a ^ p + b ^ p ≤ 1 := @@ -248,9 +258,21 @@ begin { simp [hw', fin.sum_univ_succ], }, end -end ennreal - -namespace ennreal +/-- Unweighted mean inequality, version for two elements of `ℝ≥0∞` and real exponents. -/ +theorem rpow_add_le_mul_rpow_add_rpow (z₁ z₂ : ℝ≥0∞) {p : ℝ} (hp : 1 ≤ p) : + (z₁ + z₂) ^ p ≤ 2^(p-1) * (z₁ ^ p + z₂ ^ p) := +begin + rcases eq_or_lt_of_le hp with rfl|h'p, + { simp only [rpow_one, sub_self, rpow_zero, one_mul, le_refl], }, + convert rpow_arith_mean_le_arith_mean2_rpow + (1/2) (1/2) (2 * z₁) (2 * z₂) (ennreal.add_halves 1) hp, + { simp [← mul_assoc, ennreal.inv_mul_cancel two_ne_zero two_ne_top] }, + { simp [← mul_assoc, ennreal.inv_mul_cancel two_ne_zero two_ne_top] }, + { have A : p - 1 ≠ 0 := ne_of_gt (sub_pos.2 h'p), + simp only [mul_rpow_of_nonneg _ _ (zero_le_one.trans hp), rpow_sub _ _ two_ne_zero two_ne_top, + div_eq_inv_mul, rpow_one, mul_one], + ring } +end lemma add_rpow_le_rpow_add {p : ℝ} (a b : ℝ≥0∞) (hp1 : 1 ≤ p) : a ^ p + b ^ p ≤ (a + b) ^ p := diff --git a/src/measure_theory/function/continuous_map_dense.lean b/src/measure_theory/function/continuous_map_dense.lean index 8e8820280f4cc..f0d835f76c076 100644 --- a/src/measure_theory/function/continuous_map_dense.lean +++ b/src/measure_theory/function/continuous_map_dense.lean @@ -12,7 +12,7 @@ import measure_theory.integral.bochner /-! # Approximation in Lᵖ by continuous functions -This file proves that bounded continuous functions are dense in `Lp E p μ`, for `1 ≤ p < ∞`, if the +This file proves that bounded continuous functions are dense in `Lp E p μ`, for `p < ∞`, if the domain `α` of the functions is a normal topological space and the measure `μ` is weakly regular. It also proves the same results for approximation by continuous functions with compact support when the space is locally compact and `μ` is regular. @@ -127,9 +127,9 @@ begin end /-- In a locally compact space, any function in `ℒp` can be approximated by compactly supported -continuous functions when `1 ≤ p < ∞`, version in terms of `snorm`. -/ +continuous functions when `p < ∞`, version in terms of `snorm`. -/ lemma mem_ℒp.exists_has_compact_support_snorm_sub_le - [locally_compact_space α] [μ.regular] (hp : p ≠ ∞) (h'p : 1 ≤ p) + [locally_compact_space α] [μ.regular] (hp : p ≠ ∞) {f : α → E} (hf : mem_ℒp f p μ) {ε : ℝ≥0∞} (hε : ε ≠ 0) : ∃ (g : α → E), has_compact_support g ∧ snorm (f - g) p μ ≤ ε ∧ continuous g ∧ mem_ℒp g p μ := begin @@ -140,7 +140,7 @@ begin -- It suffices to check that the set of functions we consider approximates characteristic -- functions, is stable under addition and consists of ae strongly measurable functions. -- First check the latter easy facts. - apply hf.induction_dense hp h'p _ _ _ _ hε, rotate, + apply hf.induction_dense hp _ _ _ _ hε, rotate, -- stability under addition { rintros f g ⟨f_cont, f_mem, hf⟩ ⟨g_cont, g_mem, hg⟩, exact ⟨f_cont.add g_cont, f_mem.add g_mem, hf.add hg⟩ }, @@ -150,35 +150,26 @@ begin -- We are left with approximating characteristic functions. -- This follows from `exists_continuous_snorm_sub_le_of_closed`. assume c t ht htμ ε hε, - have h'ε : ε / 2 ≠ 0, by simpa using hε, + rcases exists_Lp_half E μ p hε with ⟨δ, δpos, hδ⟩, obtain ⟨η, ηpos, hη⟩ : ∃ (η : ℝ≥0), 0 < η ∧ ∀ (s : set α), μ s ≤ η → - snorm (s.indicator (λ x, c)) p μ ≤ ε / 2, from exists_snorm_indicator_le hp c h'ε, + snorm (s.indicator (λ x, c)) p μ ≤ δ, from exists_snorm_indicator_le hp c δpos.ne', have hη_pos' : (0 : ℝ≥0∞) < η, from ennreal.coe_pos.2 ηpos, obtain ⟨s, st, s_compact, μs⟩ : ∃ s ⊆ t, is_compact s ∧ μ (t \ s) < η, from ht.exists_is_compact_diff_lt htμ.ne hη_pos'.ne', have hsμ : μ s < ∞, from (measure_mono st).trans_lt htμ, - have I1 : snorm (s.indicator (λ y, c) - t.indicator (λ y, c)) p μ ≤ ε/2, + have I1 : snorm (s.indicator (λ y, c) - t.indicator (λ y, c)) p μ ≤ δ, { rw [← snorm_neg, neg_sub, ← indicator_diff st], exact (hη _ μs.le) }, obtain ⟨k, k_compact, sk, -⟩ : ∃ (k : set α), is_compact k ∧ s ⊆ interior k ∧ k ⊆ univ, from exists_compact_between s_compact is_open_univ (subset_univ _), rcases exists_continuous_snorm_sub_le_of_closed hp s_compact.is_closed is_open_interior sk - hsμ.ne c h'ε with ⟨f, f_cont, I2, f_bound, f_support, f_mem⟩, - have I3 : snorm (f - t.indicator (λ y, c)) p μ ≤ ε, from calc - snorm (f - t.indicator (λ y, c)) p μ - = snorm ((f - s.indicator (λ y, c)) + (s.indicator (λ y, c) - t.indicator (λ y, c))) p μ : - by simp only [sub_add_sub_cancel] - ... ≤ snorm (f - s.indicator (λ y, c)) p μ - + snorm (s.indicator (λ y, c) - t.indicator (λ y, c)) p μ : - begin - refine snorm_add_le _ _ h'p, - { exact f_mem.ae_strongly_measurable.sub - (ae_strongly_measurable_const.indicator s_compact.measurable_set) }, - { exact (ae_strongly_measurable_const.indicator s_compact.measurable_set).sub - (ae_strongly_measurable_const.indicator ht) }, - end - ... ≤ ε/2 + ε/2 : add_le_add I2 I1 - ... = ε : ennreal.add_halves _, + hsμ.ne c δpos.ne' with ⟨f, f_cont, I2, f_bound, f_support, f_mem⟩, + have I3 : snorm (f - t.indicator (λ y, c)) p μ ≤ ε, + { convert (hδ _ _ (f_mem.ae_strongly_measurable.sub + (ae_strongly_measurable_const.indicator s_compact.measurable_set)) + ((ae_strongly_measurable_const.indicator s_compact.measurable_set).sub + (ae_strongly_measurable_const.indicator ht)) I2 I1).le, + simp only [sub_add_sub_cancel] }, refine ⟨f, I3, f_cont, f_mem, has_compact_support.intro k_compact (λ x hx, _)⟩, rw ← function.nmem_support, contrapose! hx, @@ -186,9 +177,9 @@ begin end /-- In a locally compact space, any function in `ℒp` can be approximated by compactly supported -continuous functions when `1 ≤ p < ∞`, version in terms of `∫`. -/ +continuous functions when `0 < p < ∞`, version in terms of `∫`. -/ lemma mem_ℒp.exists_has_compact_support_integral_rpow_sub_le - [locally_compact_space α] [μ.regular] {p : ℝ} (h'p : 1 ≤ p) + [locally_compact_space α] [μ.regular] {p : ℝ} (hp : 0 < p) {f : α → E} (hf : mem_ℒp f (ennreal.of_real p) μ) {ε : ℝ} (hε : 0 < ε) : ∃ (g : α → E), has_compact_support g ∧ ∫ x, ‖f x - g x‖^p ∂μ ≤ ε ∧ continuous g ∧ mem_ℒp g (ennreal.of_real p) μ := @@ -196,17 +187,15 @@ begin have I : 0 < ε ^ (1/p) := real.rpow_pos_of_pos hε _, have A : ennreal.of_real (ε ^ (1/p)) ≠ 0, by simp only [ne.def, ennreal.of_real_eq_zero, not_le, I], - have B : 1 ≤ ennreal.of_real p, - { convert ennreal.of_real_le_of_real h'p, exact ennreal.of_real_one.symm }, - rcases hf.exists_has_compact_support_snorm_sub_le ennreal.coe_ne_top B A + have B : ennreal.of_real p ≠ 0, by simpa only [ne.def, ennreal.of_real_eq_zero, not_le] using hp, + rcases hf.exists_has_compact_support_snorm_sub_le ennreal.coe_ne_top A with ⟨g, g_support, hg, g_cont, g_mem⟩, change snorm _ (ennreal.of_real p) _ ≤ _ at hg, refine ⟨g, g_support, _, g_cont, g_mem⟩, - rwa [(hf.sub g_mem).snorm_eq_integral_rpow_norm (zero_lt_one.trans_le B).ne' - ennreal.coe_ne_top, ennreal.of_real_le_of_real_iff I.le, one_div, - ennreal.to_real_of_real (zero_le_one.trans h'p), real.rpow_le_rpow_iff _ hε.le _] at hg, - { exact integral_nonneg (λ x, real.rpow_nonneg_of_nonneg (norm_nonneg _) _) }, - { exact inv_pos.2 (zero_lt_one.trans_le h'p) } + rwa [(hf.sub g_mem).snorm_eq_integral_rpow_norm B ennreal.coe_ne_top, + ennreal.of_real_le_of_real_iff I.le, one_div, + ennreal.to_real_of_real hp.le, real.rpow_le_rpow_iff _ hε.le (inv_pos.2 hp)] at hg, + exact integral_nonneg (λ x, real.rpow_nonneg_of_nonneg (norm_nonneg _) _), end /-- In a locally compact space, any integrable function can be approximated by compactly supported @@ -217,7 +206,7 @@ lemma integrable.exists_has_compact_support_lintegral_sub_le [locally_compact_sp ∧ integrable g μ := begin simp only [← mem_ℒp_one_iff_integrable, ← snorm_one_eq_lintegral_nnnorm] at hf ⊢, - exact hf.exists_has_compact_support_snorm_sub_le ennreal.one_ne_top le_rfl hε, + exact hf.exists_has_compact_support_snorm_sub_le ennreal.one_ne_top hε, end /-- In a locally compact space, any integrable function can be approximated by compactly supported @@ -229,12 +218,12 @@ lemma integrable.exists_has_compact_support_integral_sub_le [locally_compact_spa begin simp only [← mem_ℒp_one_iff_integrable, ← snorm_one_eq_lintegral_nnnorm, ← ennreal.of_real_one] at hf ⊢, - simpa using hf.exists_has_compact_support_integral_rpow_sub_le le_rfl hε, + simpa using hf.exists_has_compact_support_integral_rpow_sub_le zero_lt_one hε, end -/-- Any function in `ℒp` can be approximated by bounded continuous functions when `1 ≤ p < ∞`, +/-- Any function in `ℒp` can be approximated by bounded continuous functions when `p < ∞`, version in terms of `snorm`. -/ -lemma mem_ℒp.exists_bounded_continuous_snorm_sub_le [μ.weakly_regular] (hp : p ≠ ∞) (h'p : 1 ≤ p) +lemma mem_ℒp.exists_bounded_continuous_snorm_sub_le [μ.weakly_regular] (hp : p ≠ ∞) {f : α → E} (hf : mem_ℒp f p μ) {ε : ℝ≥0∞} (hε : ε ≠ 0) : ∃ (g : α →ᵇ E), snorm (f - g) p μ ≤ ε ∧ mem_ℒp g p μ := begin @@ -245,7 +234,7 @@ begin -- It suffices to check that the set of functions we consider approximates characteristic -- functions, is stable under addition and made of ae strongly measurable functions. -- First check the latter easy facts. - apply hf.induction_dense hp h'p _ _ _ _ hε, rotate, + apply hf.induction_dense hp _ _ _ _ hε, rotate, -- stability under addition { rintros f g ⟨f_cont, f_mem, f_bd⟩ ⟨g_cont, g_mem, g_bd⟩, refine ⟨f_cont.add g_cont, f_mem.add g_mem, _⟩, @@ -257,58 +246,46 @@ begin -- We are left with approximating characteristic functions. -- This follows from `exists_continuous_snorm_sub_le_of_closed`. assume c t ht htμ ε hε, - have h'ε : ε / 2 ≠ 0, by simpa using hε, + rcases exists_Lp_half E μ p hε with ⟨δ, δpos, hδ⟩, obtain ⟨η, ηpos, hη⟩ : ∃ (η : ℝ≥0), 0 < η ∧ ∀ (s : set α), μ s ≤ η → - snorm (s.indicator (λ x, c)) p μ ≤ ε / 2, from exists_snorm_indicator_le hp c h'ε, + snorm (s.indicator (λ x, c)) p μ ≤ δ, from exists_snorm_indicator_le hp c δpos.ne', have hη_pos' : (0 : ℝ≥0∞) < η, from ennreal.coe_pos.2 ηpos, obtain ⟨s, st, s_closed, μs⟩ : ∃ s ⊆ t, is_closed s ∧ μ (t \ s) < η, from ht.exists_is_closed_diff_lt htμ.ne hη_pos'.ne', have hsμ : μ s < ∞, from (measure_mono st).trans_lt htμ, - have I1 : snorm (s.indicator (λ y, c) - t.indicator (λ y, c)) p μ ≤ ε/2, + have I1 : snorm (s.indicator (λ y, c) - t.indicator (λ y, c)) p μ ≤ δ, { rw [← snorm_neg, neg_sub, ← indicator_diff st], exact (hη _ μs.le) }, rcases exists_continuous_snorm_sub_le_of_closed hp s_closed is_open_univ (subset_univ _) - hsμ.ne c h'ε with ⟨f, f_cont, I2, f_bound, -, f_mem⟩, - have I3 : snorm (f - t.indicator (λ y, c)) p μ ≤ ε, from calc - snorm (f - t.indicator (λ y, c)) p μ - = snorm ((f - s.indicator (λ y, c)) + (s.indicator (λ y, c) - t.indicator (λ y, c))) p μ : - by simp only [sub_add_sub_cancel] - ... ≤ snorm (f - s.indicator (λ y, c)) p μ - + snorm (s.indicator (λ y, c) - t.indicator (λ y, c)) p μ : - begin - refine snorm_add_le _ _ h'p, - { exact f_mem.ae_strongly_measurable.sub - (ae_strongly_measurable_const.indicator s_closed.measurable_set) }, - { exact (ae_strongly_measurable_const.indicator s_closed.measurable_set).sub - (ae_strongly_measurable_const.indicator ht) }, - end - ... ≤ ε/2 + ε/2 : add_le_add I2 I1 - ... = ε : ennreal.add_halves _, + hsμ.ne c δpos.ne' with ⟨f, f_cont, I2, f_bound, -, f_mem⟩, + have I3 : snorm (f - t.indicator (λ y, c)) p μ ≤ ε, + { convert (hδ _ _ (f_mem.ae_strongly_measurable.sub + (ae_strongly_measurable_const.indicator s_closed.measurable_set)) + ((ae_strongly_measurable_const.indicator s_closed.measurable_set).sub + (ae_strongly_measurable_const.indicator ht)) I2 I1).le, + simp only [sub_add_sub_cancel] }, refine ⟨f, I3, f_cont, f_mem, _⟩, exact (bounded_continuous_function.of_normed_add_comm_group f f_cont _ f_bound).bounded_range, end -/-- Any function in `ℒp` can be approximated by bounded continuous functions when `1 ≤ p < ∞`, +/-- Any function in `ℒp` can be approximated by bounded continuous functions when `0 < p < ∞`, version in terms of `∫`. -/ lemma mem_ℒp.exists_bounded_continuous_integral_rpow_sub_le - [μ.weakly_regular] {p : ℝ} (h'p : 1 ≤ p) + [μ.weakly_regular] {p : ℝ} (hp : 0 < p) {f : α → E} (hf : mem_ℒp f (ennreal.of_real p) μ) {ε : ℝ} (hε : 0 < ε) : ∃ (g : α →ᵇ E), ∫ x, ‖f x - g x‖^p ∂μ ≤ ε ∧ mem_ℒp g (ennreal.of_real p) μ := begin have I : 0 < ε ^ (1/p) := real.rpow_pos_of_pos hε _, have A : ennreal.of_real (ε ^ (1/p)) ≠ 0, by simp only [ne.def, ennreal.of_real_eq_zero, not_le, I], - have B : 1 ≤ ennreal.of_real p, - { convert ennreal.of_real_le_of_real h'p, exact ennreal.of_real_one.symm }, - rcases hf.exists_bounded_continuous_snorm_sub_le ennreal.coe_ne_top B A - with ⟨g, hg, g_mem⟩, + have B : ennreal.of_real p ≠ 0, by simpa only [ne.def, ennreal.of_real_eq_zero, not_le] using hp, + rcases hf.exists_bounded_continuous_snorm_sub_le ennreal.coe_ne_top A with ⟨g, hg, g_mem⟩, change snorm _ (ennreal.of_real p) _ ≤ _ at hg, refine ⟨g, _, g_mem⟩, - rwa [(hf.sub g_mem).snorm_eq_integral_rpow_norm (zero_lt_one.trans_le B).ne' - ennreal.coe_ne_top, ennreal.of_real_le_of_real_iff I.le, one_div, - ennreal.to_real_of_real (zero_le_one.trans h'p), real.rpow_le_rpow_iff _ hε.le _] at hg, - { exact integral_nonneg (λ x, real.rpow_nonneg_of_nonneg (norm_nonneg _) _) }, - { exact inv_pos.2 (zero_lt_one.trans_le h'p) } + rwa [(hf.sub g_mem).snorm_eq_integral_rpow_norm B ennreal.coe_ne_top, + ennreal.of_real_le_of_real_iff I.le, one_div, + ennreal.to_real_of_real hp.le, real.rpow_le_rpow_iff _ hε.le (inv_pos.2 hp)] at hg, + exact integral_nonneg (λ x, real.rpow_nonneg_of_nonneg (norm_nonneg _) _), end /-- Any integrable function can be approximated by bounded continuous functions, @@ -318,7 +295,7 @@ lemma integrable.exists_bounded_continuous_lintegral_sub_le [μ.weakly_regular] ∃ (g : α →ᵇ E), ∫⁻ x, ‖f x - g x‖₊ ∂μ ≤ ε ∧ integrable g μ := begin simp only [← mem_ℒp_one_iff_integrable, ← snorm_one_eq_lintegral_nnnorm] at hf ⊢, - exact hf.exists_bounded_continuous_snorm_sub_le ennreal.one_ne_top le_rfl hε, + exact hf.exists_bounded_continuous_snorm_sub_le ennreal.one_ne_top hε, end /-- Any integrable function can be approximated by bounded continuous functions, @@ -329,7 +306,7 @@ lemma integrable.exists_bounded_continuous_integral_sub_le [μ.weakly_regular] begin simp only [← mem_ℒp_one_iff_integrable, ← snorm_one_eq_lintegral_nnnorm, ← ennreal.of_real_one] at hf ⊢, - simpa using hf.exists_bounded_continuous_integral_rpow_sub_le le_rfl hε, + simpa using hf.exists_bounded_continuous_integral_rpow_sub_le zero_lt_one hε, end namespace Lp @@ -348,7 +325,7 @@ begin intros ε hε, have A : ennreal.of_real ε ≠ 0, by simp only [ne.def, ennreal.of_real_eq_zero, not_le, hε], obtain ⟨g, hg, g_mem⟩ : ∃ (g : α →ᵇ E), snorm (f - g) p μ ≤ ennreal.of_real ε ∧ mem_ℒp g p μ, - from (Lp.mem_ℒp f).exists_bounded_continuous_snorm_sub_le hp _i.out A, + from (Lp.mem_ℒp f).exists_bounded_continuous_snorm_sub_le hp A, refine ⟨g_mem.to_Lp _, _, ⟨g, rfl⟩⟩, simp only [dist_eq_norm, metric.mem_closed_ball'], rw Lp.norm_def, diff --git a/src/measure_theory/function/lp_space.lean b/src/measure_theory/function/lp_space.lean index ab45f937e1a84..851f9dc88bdf9 100644 --- a/src/measure_theory/function/lp_space.lean +++ b/src/measure_theory/function/lp_space.lean @@ -706,6 +706,20 @@ end ... ≤ snorm' f q μ + snorm' g q μ : ennreal.lintegral_Lp_add_le hf.ennnorm hg.ennnorm hq1 +lemma snorm'_add_le_of_le_one {f g : α → E} + (hf : ae_strongly_measurable f μ) (hq0 : 0 ≤ q) (hq1 : q ≤ 1) : + snorm' (f + g) q μ ≤ 2^(1/q-1) * (snorm' f q μ + snorm' g q μ) := +calc (∫⁻ a, ↑‖(f + g) a‖₊ ^ q ∂μ) ^ (1 / q) + ≤ (∫⁻ a, (((λ a, (‖f a‖₊ : ℝ≥0∞)) + + (λ a, (‖g a‖₊ : ℝ≥0∞))) a) ^ q ∂μ) ^ (1 / q) : +begin + refine ennreal.rpow_le_rpow _ (by simp [hq0] : 0 ≤ 1 / q), + refine lintegral_mono (λ a, ennreal.rpow_le_rpow _ hq0), + simp [←ennreal.coe_add, nnnorm_add_le], +end +... ≤ 2^(1/q-1) * (snorm' f q μ + snorm' g q μ) : + ennreal.lintegral_Lp_add_le_of_le_one hf.ennnorm hq0 hq1 + lemma snorm_ess_sup_add_le {f g : α → F} : snorm_ess_sup (f + g) μ ≤ snorm_ess_sup f μ + snorm_ess_sup g μ := begin @@ -729,9 +743,78 @@ begin exact snorm'_add_le hf hg hp1_real, end -lemma snorm_sub_le - {f g : α → E} (hf : ae_strongly_measurable f μ) (hg : ae_strongly_measurable g μ) (hp1 : 1 ≤ p) : - snorm (f - g) p μ ≤ snorm f p μ + snorm g p μ := +/-- A constant for the inequality `‖f + g‖_{L^p} ≤ C * (‖f‖_{L^p} + ‖g‖_{L^p})`. It is equal to `1` +for `p ≥ 1` or `p = 0`, and `2^(1/p-1)` in the more tricky interval `(0, 1)`. -/ +def Lp_add_const (p : ℝ≥0∞) : ℝ≥0∞ := +if p ∈ set.Ioo (0 : ℝ≥0∞) 1 then 2^(1/p.to_real-1) else 1 + +lemma Lp_add_const_of_one_le {p : ℝ≥0∞} (hp : 1 ≤ p) : Lp_add_const p = 1 := +begin + rw [Lp_add_const, if_neg], + assume h, + exact lt_irrefl _ (h.2.trans_le hp), +end + +lemma Lp_add_const_zero : Lp_add_const 0 = 1 := +begin + rw [Lp_add_const, if_neg], + assume h, + exact lt_irrefl _ (h.1), +end + +lemma Lp_add_const_lt_top (p : ℝ≥0∞) : Lp_add_const p < ∞ := +begin + rw [Lp_add_const], + split_ifs, + { apply ennreal.rpow_lt_top_of_nonneg _ ennreal.two_ne_top, + simp only [one_div, sub_nonneg], + apply one_le_inv (ennreal.to_real_pos h.1.ne' (h.2.trans ennreal.one_lt_top).ne), + simpa using ennreal.to_real_mono ennreal.one_ne_top h.2.le }, + { exact ennreal.one_lt_top } +end + +lemma snorm_add_le' + {f g : α → E} (hf : ae_strongly_measurable f μ) (hg : ae_strongly_measurable g μ) (p : ℝ≥0∞) : + snorm (f + g) p μ ≤ Lp_add_const p * (snorm f p μ + snorm g p μ) := +begin + rcases eq_or_ne p 0 with rfl|hp, + { simp only [snorm_exponent_zero, add_zero, mul_zero, le_zero_iff] }, + rcases lt_or_le p 1 with h'p|h'p, + { simp only [snorm_eq_snorm' hp (h'p.trans ennreal.one_lt_top).ne], + convert snorm'_add_le_of_le_one hf ennreal.to_real_nonneg _, + { have : p ∈ set.Ioo (0 : ℝ≥0∞) 1 := ⟨hp.bot_lt, h'p⟩, + simp only [Lp_add_const, if_pos this] }, + { simpa using ennreal.to_real_mono ennreal.one_ne_top h'p.le } }, + { simp [Lp_add_const_of_one_le h'p], + exact snorm_add_le hf hg h'p } +end + +variables (μ E) +/-- Technical lemma to control the addition of functions in `L^p` even for `p < 1`: Given `δ > 0`, +there exists `η` such that two functions bounded by `η` in `L^p` have a sum bounded by `δ`. One +could take `η = δ / 2` for `p ≥ 1`, but the point of the lemma is that it works also for `p < 1`. +-/ +lemma exists_Lp_half (p : ℝ≥0∞) {δ : ℝ≥0∞} (hδ : δ ≠ 0) : + ∃ (η : ℝ≥0∞), 0 < η ∧ ∀ (f g : α → E) (hf : ae_strongly_measurable f μ) + (hg : ae_strongly_measurable g μ) (Hf : snorm f p μ ≤ η) (Hg : snorm g p μ ≤ η), + snorm (f + g) p μ < δ := +begin + have : tendsto (λ (η : ℝ≥0∞), Lp_add_const p * (η + η)) (𝓝[>] 0) (𝓝 (Lp_add_const p * (0 + 0))), + from (ennreal.tendsto.const_mul (tendsto_id.add tendsto_id) + (or.inr (Lp_add_const_lt_top p).ne)).mono_left nhds_within_le_nhds, + simp only [add_zero, mul_zero] at this, + rcases (((tendsto_order.1 this).2 δ hδ.bot_lt).and self_mem_nhds_within).exists + with ⟨η, hη, ηpos⟩, + refine ⟨η, ηpos, λ f g hf hg Hf Hg, _⟩, + calc snorm (f + g) p μ ≤ Lp_add_const p * (snorm f p μ + snorm g p μ) : snorm_add_le' hf hg p + ... ≤ Lp_add_const p * (η + η) : mul_le_mul_of_nonneg_left (add_le_add Hf Hg) bot_le + ... < δ : hη +end +variables {μ E} + +lemma snorm_sub_le' + {f g : α → E} (hf : ae_strongly_measurable f μ) (hg : ae_strongly_measurable g μ) (p : ℝ≥0∞) : + snorm (f - g) p μ ≤ Lp_add_const p * (snorm f p μ + snorm g p μ) := calc snorm (f - g) p μ = snorm (f + - g) p μ : by rw sub_eq_add_neg -- We cannot use snorm_add_le on f and (-g) because we don't have `ae_measurable (-g) μ`, since -- we don't suppose `[borel_space E]`. @@ -739,55 +822,23 @@ calc snorm (f - g) p μ = snorm (f + - g) p μ : by rw sub_eq_add_neg ... ≤ snorm (λ x, ‖f x‖ + ‖- g x‖) p μ : by { refine snorm_mono_real (λ x, _), rw norm_norm, exact norm_add_le _ _, } ... = snorm (λ x, ‖f x‖ + ‖g x‖) p μ : by simp_rw norm_neg -... ≤ snorm (λ x, ‖f x‖) p μ + snorm (λ x, ‖g x‖) p μ : snorm_add_le hf.norm hg.norm hp1 -... = snorm f p μ + snorm g p μ : by rw [← snorm_norm f, ← snorm_norm g] - -lemma snorm_add_lt_top_of_one_le {f g : α → E} (hf : mem_ℒp f p μ) (hg : mem_ℒp g p μ) - (hq1 : 1 ≤ p) : snorm (f + g) p μ < ∞ := -lt_of_le_of_lt (snorm_add_le hf.1 hg.1 hq1) (ennreal.add_lt_top.mpr ⟨hf.2, hg.2⟩) +... ≤ Lp_add_const p * (snorm (λ x, ‖f x‖) p μ + snorm (λ x, ‖g x‖) p μ) : + snorm_add_le' hf.norm hg.norm p +... = Lp_add_const p * (snorm f p μ + snorm g p μ) : by rw [← snorm_norm f, ← snorm_norm g] -lemma snorm'_add_lt_top_of_le_one - {f g : α → E} (hf : ae_strongly_measurable f μ) - (hf_snorm : snorm' f q μ < ∞) (hg_snorm : snorm' g q μ < ∞) (hq_pos : 0 < q) (hq1 : q ≤ 1) : - snorm' (f + g) q μ < ∞ := -calc (∫⁻ a, ↑‖(f + g) a‖₊ ^ q ∂μ) ^ (1 / q) - ≤ (∫⁻ a, (((λ a, (‖f a‖₊ : ℝ≥0∞)) - + (λ a, (‖g a‖₊ : ℝ≥0∞))) a) ^ q ∂μ) ^ (1 / q) : -begin - refine ennreal.rpow_le_rpow _ (by simp [hq_pos.le] : 0 ≤ 1 / q), - refine lintegral_mono (λ a, ennreal.rpow_le_rpow _ hq_pos.le), - simp [←ennreal.coe_add, nnnorm_add_le], -end -... ≤ (∫⁻ a, (‖f a‖₊ : ℝ≥0∞) ^ q + (‖g a‖₊ : ℝ≥0∞) ^ q ∂μ) ^ (1 / q) : -begin - refine ennreal.rpow_le_rpow (lintegral_mono (λ a, _)) (by simp [hq_pos.le] : 0 ≤ 1 / q), - exact ennreal.rpow_add_le_add_rpow _ _ hq_pos.le hq1, -end -... < ∞ : -begin - refine ennreal.rpow_lt_top_of_nonneg (by simp [hq_pos.le] : 0 ≤ 1 / q) _, - rw [lintegral_add_left' (hf.ennnorm.pow_const q), ennreal.add_ne_top], - exact ⟨(lintegral_rpow_nnnorm_lt_top_of_snorm'_lt_top hq_pos hf_snorm).ne, - (lintegral_rpow_nnnorm_lt_top_of_snorm'_lt_top hq_pos hg_snorm).ne⟩, -end +lemma snorm_sub_le + {f g : α → E} (hf : ae_strongly_measurable f μ) (hg : ae_strongly_measurable g μ) (hp : 1 ≤ p) : + snorm (f - g) p μ ≤ snorm f p μ + snorm g p μ := +by simpa [Lp_add_const_of_one_le hp] using snorm_sub_le' hf hg p lemma snorm_add_lt_top {f g : α → E} (hf : mem_ℒp f p μ) (hg : mem_ℒp g p μ) : - snorm (f + g) p μ < ∞ := + snorm (f + g) p μ < ∞ := calc +snorm (f + g) p μ ≤ Lp_add_const p * (snorm f p μ + snorm g p μ) : + snorm_add_le' hf.ae_strongly_measurable hg.ae_strongly_measurable p +... < ∞ : begin - by_cases h0 : p = 0, - { simp [h0], }, - rw ←ne.def at h0, - cases le_total 1 p with hp1 hp1, - { exact snorm_add_lt_top_of_one_le hf hg hp1, }, - have hp_top : p ≠ ∞, from (lt_of_le_of_lt hp1 ennreal.coe_lt_top).ne, - have hp_pos : 0 < p.to_real, - { rw [← ennreal.zero_to_real, @ennreal.to_real_lt_to_real 0 p ennreal.coe_ne_top hp_top], - exact ((zero_le p).lt_of_ne h0.symm), }, - have hp1_real : p.to_real ≤ 1, - { rwa [← ennreal.one_to_real, @ennreal.to_real_le_to_real p 1 hp_top ennreal.coe_ne_top], }, - rw snorm_eq_snorm' h0 hp_top, - rw [mem_ℒp, snorm_eq_snorm' h0 hp_top] at hf hg, - exact snorm'_add_lt_top_of_le_one hf.1 hf.2 hg.2 hp_pos hp1_real, + apply ennreal.mul_lt_top (Lp_add_const_lt_top p).ne, + exact ((ennreal.add_lt_top).2 ⟨hf.2, hg.2⟩).ne, end section map_measure @@ -918,9 +969,6 @@ begin simp [snorm_eq_snorm' h0 h_top], end -section borel_space --- variable [borel_space E] - lemma mem_ℒp.neg {f : α → E} (hf : mem_ℒp f p μ) : mem_ℒp (-f) p μ := ⟨ae_strongly_measurable.neg hf.1, by simp [hf.right]⟩ @@ -1160,8 +1208,6 @@ end end has_measurable_add -end borel_space - section normed_space variables {𝕜 : Type*} [normed_field 𝕜] [normed_space 𝕜 E] [normed_space 𝕜 F] diff --git a/src/measure_theory/function/simple_func_dense_lp.lean b/src/measure_theory/function/simple_func_dense_lp.lean index b8a92f376fba8..1853843cf77f6 100644 --- a/src/measure_theory/function/simple_func_dense_lp.lean +++ b/src/measure_theory/function/simple_func_dense_lp.lean @@ -186,6 +186,31 @@ lemma tendsto_approx_on_range_Lp [borel_space E] by simpa only [Lp.tendsto_Lp_iff_tendsto_ℒp''] using tendsto_approx_on_range_Lp_snorm hp_ne_top fmeas hf.2 +/-- Any function in `ℒp` can be approximated by a simple function if `p < ∞`. -/ +lemma _root_.measure_theory.mem_ℒp.exists_simple_func_snorm_sub_lt + {E : Type*} [normed_add_comm_group E] + {f : β → E} {μ : measure β} (hf : mem_ℒp f p μ) (hp_ne_top : p ≠ ∞) {ε : ℝ≥0∞} (hε : ε ≠ 0) : + ∃ (g : β →ₛ E), snorm (f - g) p μ < ε ∧ mem_ℒp g p μ := +begin + borelize E, + let f' := hf.1.mk f, + suffices H : ∃ (g : β →ₛ E), snorm (f' - g) p μ < ε ∧ mem_ℒp g p μ, + { rcases H with ⟨g, hg, g_mem⟩, + refine ⟨g, _, g_mem⟩, + convert hg using 1, + apply snorm_congr_ae, + filter_upwards [hf.1.ae_eq_mk] with x hx, + simpa only [pi.sub_apply, sub_left_inj] using hx }, + have hf' : mem_ℒp f' p μ, from hf.ae_eq hf.1.ae_eq_mk, + have f'meas : measurable f' := hf.1.measurable_mk, + haveI : separable_space (range f' ∪ {0} : set E), + from strongly_measurable.separable_space_range_union_singleton hf.1.strongly_measurable_mk, + rcases ((tendsto_order.1 (tendsto_approx_on_range_Lp_snorm hp_ne_top f'meas hf'.2)).2 + ε hε.bot_lt).exists with ⟨n, hn⟩, + rw [← snorm_neg, neg_sub] at hn, + exact ⟨_, hn, mem_ℒp_approx_on_range f'meas hf' _⟩, +end + end Lp /-! ### L1 approximation by simple functions -/ @@ -897,7 +922,7 @@ end /-- If a set of ae strongly measurable functions is stable under addition and approximates characteristic functions in `ℒp`, then it is dense in `ℒp`. -/ -lemma mem_ℒp.induction_dense (hp_ne_top : p ≠ ∞) (h'p : 1 ≤ p) (P : (α → E) → Prop) +lemma mem_ℒp.induction_dense (hp_ne_top : p ≠ ∞) (P : (α → E) → Prop) (h0P : ∀ (c : E) ⦃s : set α⦄, measurable_set s → μ s < ∞ → ∀ {ε : ℝ≥0∞}, ε ≠ 0 → (∃ (g : α → E), snorm (g - s.indicator (λ x, c)) p μ ≤ ε ∧ P g)) (h1P : ∀ f g, P f → P g → P (f + g)) @@ -905,48 +930,45 @@ lemma mem_ℒp.induction_dense (hp_ne_top : p ≠ ∞) (h'p : 1 ≤ p) (P : (α {f : α → E} (hf : mem_ℒp f p μ) {ε : ℝ≥0∞} (hε : ε ≠ 0) : ∃ (g : α → E), snorm (f - g) p μ ≤ ε ∧ P g := begin - haveI : fact (1 ≤ p) := ⟨h'p⟩, - revert f hf ε, - refine mem_ℒp.induction hp_ne_top _ _ _ _ _, - { assume c s hs hμs ε εpos, - rcases h0P c hs hμs εpos with ⟨g, hg, Pg⟩, - rw [← snorm_neg, neg_sub] at hg, - exact ⟨g, hg, Pg⟩ }, - { assume f f' hff' hf hf' Hf Hf' ε εpos, - have A : ε / 2 ≠ 0, by simp [εpos], - rcases Hf A with ⟨g, hfg, Pg⟩, - rcases Hf' A with ⟨g', hf'g', Pg'⟩, - refine ⟨g + g', _, h1P g g' Pg Pg'⟩, - calc snorm (f + f' - (g + g')) p μ - = snorm ((f - g) + (f' - g')) p μ : by { congr' 1, abel } - ... ≤ snorm (f - g) p μ + snorm (f' - g') p μ : - snorm_add_le (hf.ae_strongly_measurable.sub (h2P g Pg)) - (hf'.ae_strongly_measurable.sub (h2P g' Pg')) h'p - ... ≤ ε / 2 + ε / 2 : add_le_add hfg hf'g' - ... = ε : ennreal.add_halves _ }, - { rw is_closed_iff_nhds, - assume f hf ε εpos, - have A : ε / 2 ≠ 0, by simp [εpos], - rcases hf (emetric.ball f (ε/2)) (emetric.ball_mem_nhds _ A.bot_lt) with ⟨f', hf', h'f'⟩, - rcases h'f' A with ⟨g, hg, Pg⟩, - refine ⟨g, _, Pg⟩, - calc snorm (f - g) p μ = snorm ((f - f') + (f' - g)) p μ : by simp only [sub_add_sub_cancel] - ... ≤ snorm (f - f') p μ + snorm (f' - g) p μ : - snorm_add_le ((Lp.mem_ℒp f).sub (Lp.mem_ℒp f')).ae_strongly_measurable - ((Lp.mem_ℒp f').ae_strongly_measurable.sub (h2P g Pg)) h'p - ... ≤ ε / 2 + ε / 2 : - begin - refine add_le_add _ hg, - rw [← snorm_neg, neg_sub], - simp only [Lp.edist_def, emetric.mem_ball] at hf', - exact hf'.le - end - ... = ε : ennreal.add_halves _ }, - { assume f f' hff' hf Hf ε εpos, - rcases Hf εpos with ⟨g, hg, Pg⟩, + rcases eq_or_ne p 0 with rfl|hp_pos, + { rcases h0P (0 : E) measurable_set.empty + (by simp only [measure_empty, with_top.zero_lt_top]) hε with ⟨g, hg, Pg⟩, + exact ⟨g, by simp only [snorm_exponent_zero, zero_le'], Pg⟩ }, + suffices H : ∀ (f' : α →ₛ E) (δ : ℝ≥0∞) (hδ : δ ≠ 0), mem_ℒp f' p μ → + ∃ g, snorm (f' - g) p μ ≤ δ ∧ P g, + { obtain ⟨η, ηpos, hη⟩ := exists_Lp_half E μ p hε, + rcases hf.exists_simple_func_snorm_sub_lt hp_ne_top ηpos.ne' with ⟨f', hf', f'_mem⟩, + rcases H f' η ηpos.ne' f'_mem with ⟨g, hg, Pg⟩, refine ⟨g, _, Pg⟩, - have : f - g =ᵐ[μ] f' - g := hff'.sub (filter.germ.coe_eq.mp rfl), - rwa ← snorm_congr_ae this } + convert (hη _ _ (hf.ae_strongly_measurable.sub f'.ae_strongly_measurable) + (f'.ae_strongly_measurable.sub (h2P g Pg)) hf'.le hg).le, + simp only [sub_add_sub_cancel] }, + refine simple_func.induction _ _, + { intros c s hs ε εpos Hs, + rcases eq_or_ne c 0 with rfl|hc, + { rcases h0P (0 : E) measurable_set.empty + (by simp only [measure_empty, with_top.zero_lt_top]) εpos with ⟨g, hg, Pg⟩, + rw [← snorm_neg, neg_sub] at hg, + refine ⟨g, _, Pg⟩, + convert hg, + ext x, + simp only [simple_func.const_zero, simple_func.coe_piecewise, simple_func.coe_zero, + piecewise_eq_indicator, indicator_zero', pi.zero_apply, indicator_zero] }, + { have : μ s < ∞, + from (simple_func.measure_lt_top_of_mem_ℒp_indicator hp_pos hp_ne_top hc hs Hs), + rcases h0P c hs this εpos with ⟨g, hg, Pg⟩, + rw [← snorm_neg, neg_sub] at hg, + exact ⟨g, hg, Pg⟩ } }, + { intros f f' hff' hf hf' δ δpos int_ff', + obtain ⟨η, ηpos, hη⟩ := exists_Lp_half E μ p δpos, + rw [simple_func.coe_add, + mem_ℒp_add_of_disjoint hff' f.strongly_measurable f'.strongly_measurable] at int_ff', + rcases hf η ηpos.ne' int_ff'.1 with ⟨g, hg, Pg⟩, + rcases hf' η ηpos.ne' int_ff'.2 with ⟨g', hg', Pg'⟩, + refine ⟨g + g', _, h1P g g' Pg Pg'⟩, + convert (hη _ _ (f.ae_strongly_measurable.sub (h2P g Pg)) + (f'.ae_strongly_measurable.sub (h2P g' Pg')) hg hg').le, + abel } end section integrable diff --git a/src/measure_theory/integral/mean_inequalities.lean b/src/measure_theory/integral/mean_inequalities.lean index 9c902c4f1671a..9c8942b2eba36 100644 --- a/src/measure_theory/integral/mean_inequalities.lean +++ b/src/measure_theory/integral/mean_inequalities.lean @@ -377,6 +377,29 @@ begin exact lintegral_Lp_add_le_aux hpq hf hf_top hg hg_top h0 htop, end +/-- Variant of Minkowski's inequality for functions `α → ℝ≥0∞` in `ℒp` with `p ≤ 1`: the `ℒp` +seminorm of the sum of two functions is bounded by a constant multiple of the sum +of their `ℒp` seminorms. -/ +theorem lintegral_Lp_add_le_of_le_one {p : ℝ} {f g : α → ℝ≥0∞} + (hf : ae_measurable f μ) (hp0 : 0 ≤ p) (hp1 : p ≤ 1) : + (∫⁻ a, ((f + g) a)^p ∂ μ) ^ (1/p) ≤ + 2^(1/p-1) * ((∫⁻ a, (f a)^p ∂μ) ^ (1/p) + (∫⁻ a, (g a)^p ∂μ) ^ (1/p)) := +begin + rcases eq_or_lt_of_le hp0 with rfl|hp, + { simp only [pi.add_apply, rpow_zero, lintegral_one, _root_.div_zero, zero_sub], + norm_num, + rw [rpow_neg, rpow_one, ennreal.inv_mul_cancel two_ne_zero two_ne_top], + exact le_rfl }, + calc (∫⁻ a, (f + g) a ^ p ∂μ) ^ (1 / p) ≤ (∫⁻ a, (f a)^p ∂ μ + ∫⁻ a, (g a)^p ∂ μ) ^ (1/p) : + begin + apply rpow_le_rpow _ (div_nonneg zero_le_one hp0), + rw ← lintegral_add_left' (hf.pow_const p), + exact lintegral_mono (λ a, rpow_add_le_add_rpow _ _ hp0 hp1) + end + ... ≤ 2 ^ (1/p-1) * ((∫⁻ a, f a ^ p ∂μ) ^ (1/p) + (∫⁻ a, g a ^ p ∂μ) ^ (1/p)) : + rpow_add_le_mul_rpow_add_rpow _ _ ((one_le_div hp).2 hp1) +end + end ennreal /-- Hölder's inequality for functions `α → ℝ≥0`. The integral of the product of two functions From 49b7f94aab3a3bdca1f9f34c5d818afb253b3993 Mon Sep 17 00:00:00 2001 From: Felix-Weilacher Date: Mon, 17 Apr 2023 20:49:23 +0000 Subject: [PATCH 0836/1291] feat(topology/perfect): Schemes, embedding of the Cantor space (#18248) Add to `topology.perfect` the following theorem: In a complete metric space, any nonempty perfect set admits a continuous embedding of the Cantor space. The proof uses an object which in descriptive set theory is sometimes called a "scheme". Some attempt was made to include in `topology.metric_space.cantor_scheme` a general theory of these schemes, since they should be useful down the line in other results as well. We also define `pi_nat.res` in `topology.metric_space.pi_nat`. --- src/topology/metric_space/cantor_scheme.lean | 195 +++++++++++++++++++ src/topology/metric_space/pi_nat.lean | 65 ++++++- src/topology/perfect.lean | 115 ++++++++++- 3 files changed, 370 insertions(+), 5 deletions(-) create mode 100644 src/topology/metric_space/cantor_scheme.lean diff --git a/src/topology/metric_space/cantor_scheme.lean b/src/topology/metric_space/cantor_scheme.lean new file mode 100644 index 0000000000000..276a911ae5514 --- /dev/null +++ b/src/topology/metric_space/cantor_scheme.lean @@ -0,0 +1,195 @@ +/- +Copyright (c) 2023 Felix Weilacher. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Felix Weilacher +-/ +import topology.metric_space.pi_nat + +/-! +# (Topological) Schemes and their induced maps + +In topology, and especially descriptive set theory, one often constructs functions `(ℕ → β) → α`, +where α is some topological space and β is a discrete space, as an appropriate limit of some map +`list β → set α`. We call the latter type of map a "`β`-scheme on `α`". + +This file develops the basic, abstract theory of these schemes and the functions they induce. + +## Main Definitions + +* `cantor_scheme.induced_map A` : The aforementioned "limit" of a scheme `A : list β → set α`. + This is a partial function from `ℕ → β` to `a`, + implemented here as an object of type `Σ s : set (ℕ → β), s → α`. + That is, `(induced_map A).1` is the domain and `(induced_map A).2` is the function. + +## Implementation Notes + +We consider end-appending to be the fundamental way to build lists (say on `β`) inductively, +as this interacts better with the topology on `ℕ → β`. +As a result, functions like `list.nth` or `stream.take` do not have their intended meaning +in this file. See instead `pi_nat.res`. + +## References + +* [kechris1995] (Chapters 6-7) + +## Tags + +scheme, cantor scheme, lusin scheme, approximation. + +-/ + +namespace cantor_scheme + +open list function filter set pi_nat +open_locale classical topology + +variables {β α : Type*} (A : list β → set α) + +/-- From a `β`-scheme on `α` `A`, we define a partial function from `(ℕ → β)` to `α` +which sends each infinite sequence `x` to an element of the intersection along the +branch corresponding to `x`, if it exists. +We call this the map induced by the scheme. -/ +noncomputable def induced_map : Σ s : set (ℕ → β), s → α := +⟨λ x, set.nonempty ⋂ n : ℕ, A (res x n), λ x, x.property.some⟩ + +section topology + +/-- A scheme is antitone if each set contains its children. -/ +protected def antitone : Prop := ∀ l : list β, ∀ a : β, A (a :: l) ⊆ A l + +/-- A useful strengthening of being antitone is to require that each set contains +the closure of each of its children. -/ +def closure_antitone [topological_space α] : Prop := +∀ l : list β, ∀ a : β, closure (A (a :: l)) ⊆ A l + +/-- A scheme is disjoint if the children of each set of pairwise disjoint. -/ +protected def disjoint : Prop := +∀ l : list β, _root_.pairwise $ λ a b, disjoint (A (a :: l)) (A (b :: l)) + +variable {A} + +/-- If `x` is in the domain of the induced map of a scheme `A`, +its image under this map is in each set along the corresponding branch. -/ +lemma map_mem (x : (induced_map A).1) (n : ℕ) : + (induced_map A).2 x ∈ A (res x n) := +begin + have := x.property.some_mem, + rw mem_Inter at this, + exact this n, +end + +protected lemma closure_antitone.antitone [topological_space α] (hA : closure_antitone A) : + cantor_scheme.antitone A := +λ l a, subset_closure.trans (hA l a) + +protected lemma antitone.closure_antitone [topological_space α] (hanti : cantor_scheme.antitone A) + (hclosed : ∀ l, is_closed (A l)) : closure_antitone A := +λ l a, (hclosed _).closure_eq.subset.trans (hanti _ _) + +/-- A scheme where the children of each set are pairwise disjoint induces an injective map. -/ +theorem disjoint.map_injective (hA : cantor_scheme.disjoint A) : injective (induced_map A).2 := +begin + rintros ⟨x, hx⟩ ⟨y, hy⟩ hxy, + refine subtype.coe_injective (res_injective _), + dsimp, + ext n : 1, + induction n with n ih, { simp }, + simp only [res_succ], + refine ⟨_, ih⟩, + contrapose hA, + simp only [cantor_scheme.disjoint, _root_.pairwise, ne.def, not_forall, exists_prop], + refine ⟨res x n, _, _, hA, _⟩, + rw not_disjoint_iff, + refine ⟨(induced_map A).2 ⟨x, hx⟩, _, _⟩, + { rw ← res_succ, + apply map_mem, }, + rw [hxy, ih, ← res_succ], + apply map_mem, +end + +end topology + +section metric + +variable [pseudo_metric_space α] + +variable (A) + +/-- A scheme on a metric space has vanishing diameter if diameter approaches 0 along each branch. -/ +def vanishing_diam : Prop := +∀ x : ℕ → β, tendsto (λ n : ℕ, emetric.diam (A (res x n))) at_top (𝓝 0) + +variable {A} + +lemma vanishing_diam.dist_lt (hA : vanishing_diam A) (ε : ℝ) (ε_pos : 0 < ε) (x : ℕ → β) : + ∃ n : ℕ, ∀ y z ∈ A (res x n), dist y z < ε := +begin + specialize hA x, + rw ennreal.tendsto_at_top_zero at hA, + cases hA (ennreal.of_real (ε / 2)) + (by { simp only [gt_iff_lt, ennreal.of_real_pos], linarith }) with n hn, + use n, + intros y hy z hz, + rw [← ennreal.of_real_lt_of_real_iff ε_pos, ← edist_dist], + apply lt_of_le_of_lt (emetric.edist_le_diam_of_mem hy hz), + apply lt_of_le_of_lt (hn _ (le_refl _)), + rw ennreal.of_real_lt_of_real_iff ε_pos, + linarith, +end + +/-- A scheme with vanishing diameter along each branch induces a continuous map. -/ +theorem vanishing_diam.map_continuous [topological_space β] [discrete_topology β] + (hA : vanishing_diam A) : continuous (induced_map A).2 := +begin + rw metric.continuous_iff', + rintros ⟨x, hx⟩ ε ε_pos, + cases hA.dist_lt _ ε_pos x with n hn, + rw _root_.eventually_nhds_iff, + refine ⟨coe ⁻¹' cylinder x n, _, _, by simp⟩, + { rintros ⟨y, hy⟩ hyx, + rw [mem_preimage, subtype.coe_mk, cylinder_eq_res, mem_set_of] at hyx, + apply hn, + { rw ← hyx, + apply map_mem, }, + apply map_mem, }, + apply continuous_subtype_coe.is_open_preimage, + apply is_open_cylinder, +end + +/-- A scheme on a complete space with vanishing diameter +such that each set contains the closure of its children +induces a total map. -/ +theorem closure_antitone.map_of_vanishing_diam [complete_space α] + (hdiam : vanishing_diam A) (hanti : closure_antitone A) (hnonempty : ∀ l, (A l).nonempty) : + (induced_map A).1 = univ := +begin + rw eq_univ_iff_forall, + intro x, + choose u hu using λ n, hnonempty (res x n), + have umem : ∀ n m : ℕ, n ≤ m → u m ∈ A (res x n), + { have : antitone (λ n : ℕ, A (res x n)), + { refine antitone_nat_of_succ_le _, + intro n, + apply hanti.antitone, }, + intros n m hnm, + exact this hnm (hu _), }, + have : cauchy_seq u, + { rw metric.cauchy_seq_iff, + intros ε ε_pos, + cases hdiam.dist_lt _ ε_pos x with n hn, + use n, + intros m₀ hm₀ m₁ hm₁, + apply hn; apply umem; assumption, }, + cases cauchy_seq_tendsto_of_complete this with y hy, + use y, + rw mem_Inter, + intro n, + apply hanti _ (x n), + apply mem_closure_of_tendsto hy, + rw eventually_at_top, + exact ⟨n.succ, umem _⟩, +end + +end metric + +end cantor_scheme diff --git a/src/topology/metric_space/pi_nat.lean b/src/topology/metric_space/pi_nat.lean index b3e2176c72d63..5311742e8272f 100644 --- a/src/topology/metric_space/pi_nat.lean +++ b/src/topology/metric_space/pi_nat.lean @@ -196,6 +196,60 @@ lemma update_mem_cylinder (x : Π n, E n) (n : ℕ) (y : E n) : update x n y ∈ cylinder x n := mem_cylinder_iff.2 (λ i hi, by simp [hi.ne]) +section res + +variable {α : Type*} +open list + +/-- In the case where `E` has constant value `α`, +the cylinder `cylinder x n` can be identified with the element of `list α` +consisting of the first `n` entries of `x`. See `cylinder_eq_res`. +We call this list `res x n`, the restriction of `x` to `n`.-/ +def res (x : ℕ → α) : ℕ → list α +| 0 := nil +| (nat.succ n) := x n :: res n + +@[simp] lemma res_zero (x : ℕ → α) : res x 0 = @nil α := rfl +@[simp] lemma res_succ (x : ℕ → α) (n : ℕ) : res x n.succ = x n :: res x n := rfl + +@[simp] lemma res_length (x : ℕ → α) (n : ℕ) : (res x n).length = n := +by induction n; simp [*] + +/-- The restrictions of `x` and `y` to `n` are equal if and only if `x m = y m` for all `m < n`.-/ +lemma res_eq_res {x y : ℕ → α} {n : ℕ} : res x n = res y n ↔ ∀ ⦃m⦄, m < n → x m = y m := +begin + split; intro h; induction n with n ih, { simp }, + { intros m hm, + rw nat.lt_succ_iff_lt_or_eq at hm, + simp only [res_succ] at h, + cases hm with hm hm, + { exact ih h.2 hm }, + rw hm, + exact h.1, }, + { simp }, + simp only [res_succ], + refine ⟨h (nat.lt_succ_self _), ih (λ m hm, _)⟩, + exact h (hm.trans (nat.lt_succ_self _)), +end + +lemma res_injective : injective (@res α) := +begin + intros x y h, + ext n, + apply (res_eq_res).mp _ (nat.lt_succ_self _), + rw h, +end + +/-- `cylinder x n` is equal to the set of sequences `y` with the same restriction to `n` as `x`.-/ +theorem cylinder_eq_res (x : ℕ → α) (n : ℕ) : cylinder x n = {y | res y n = res x n} := +begin + ext y, + dsimp [cylinder], + rw res_eq_res, +end + +end res + /-! ### A distance function on `Π n, E n` @@ -305,13 +359,18 @@ end variables (E) [∀ n, topological_space (E n)] [∀ n, discrete_topology (E n)] -lemma is_topological_basis_cylinders : +lemma is_open_cylinder (x : Π n, E n) (n : ℕ) : is_open (cylinder x n) := +begin + rw pi_nat.cylinder_eq_pi, + exact is_open_set_pi (finset.range n).finite_to_set (λ a ha, is_open_discrete _), +end + +lemma is_topological_basis_cylinders : is_topological_basis {s : set (Π n, E n) | ∃ (x : Π n, E n) (n : ℕ), s = cylinder x n} := begin apply is_topological_basis_of_open_of_nhds, { rintros u ⟨x, n, rfl⟩, - rw cylinder_eq_pi, - exact is_open_set_pi (finset.range n).finite_to_set (λ a ha, is_open_discrete _) }, + apply is_open_cylinder, }, { assume x u hx u_open, obtain ⟨v, ⟨U, F, hUF, rfl⟩, xU, Uu⟩ : ∃ (v : set (Π (i : ℕ), E i)) (H : v ∈ {S : set (Π (i : ℕ), E i) | ∃ (U : Π (i : ℕ), set (E i)) (F : finset ℕ), diff --git a/src/topology/perfect.lean b/src/topology/perfect.lean index 24609fd926296..a92dd35cf4e0d 100644 --- a/src/topology/perfect.lean +++ b/src/topology/perfect.lean @@ -5,6 +5,7 @@ Authors: Felix Weilacher -/ import topology.separation import topology.bases +import topology.metric_space.cantor_scheme /-! # Perfect Sets @@ -19,6 +20,8 @@ including a version of the Cantor-Bendixson Theorem. * `perfect C`: A set `C` is perfect, meaning it is closed and every point of it is an accumulation point of itself. +* `set.scheme β α`: A `β`-scheme on `α`, a collection of subsets of `α` indexed by `list β`. + Used to construct maps `(β → ℕ) → α` as limiting objects. ## Main Statements @@ -28,6 +31,8 @@ including a version of the Cantor-Bendixson Theorem. * `exists_countable_union_perfect_of_is_closed`: One version of the **Cantor-Bendixson Theorem**: A closed set in a second countable space can be written as the union of a countable set and a perfect set. +* `exists_nat_bool_injection_of_perfect_nonempty`: A perfect nonempty set in a complete metric space + admits an embedding from the Cantor space. ## Implementation Notes @@ -39,17 +44,19 @@ see `preperfect_iff_perfect_closure`. ## References -* [kechris1995] (Chapter 6) +* [kechris1995] (Chapters 6-7) ## Tags -accumulation point, perfect set, Cantor-Bendixson. +accumulation point, perfect set, cantor-bendixson. -/ open_locale topology filter open topological_space filter set +section basic + variables {α : Type*} [topological_space α] {C : set α} /-- If `x` is an accumulation point of a set `C` and `U` is a neighborhood of `x`, @@ -214,3 +221,107 @@ begin end end kernel +end basic + +section cantor_inj + +open function +open_locale ennreal +variables {α : Type*} [metric_space α] {C : set α} (hC : perfect C) {ε : ℝ≥0∞} +include hC + +private lemma perfect.small_diam_aux (ε_pos : 0 < ε) {x : α} (xC : x ∈ C) : + let D := closure (emetric.ball x (ε / 2) ∩ C) in + perfect D ∧ D.nonempty ∧ D ⊆ C ∧ emetric.diam D ≤ ε := +begin + have : x ∈ (emetric.ball x (ε / 2)), + { apply emetric.mem_ball_self, + rw ennreal.div_pos_iff, + exact ⟨ne_of_gt ε_pos, by norm_num⟩, }, + have := hC.closure_nhds_inter x xC this emetric.is_open_ball, + refine ⟨this.1, this.2, _, _⟩, + { rw is_closed.closure_subset_iff hC.closed, + apply inter_subset_right, }, + rw emetric.diam_closure, + apply le_trans (emetric.diam_mono (inter_subset_left _ _)), + convert emetric.diam_ball, + rw [mul_comm, ennreal.div_mul_cancel]; norm_num, +end + +variable (hnonempty : C.nonempty) +include hnonempty + +/-- A refinement of `perfect.splitting` for metric spaces, where we also control +the diameter of the new perfect sets. -/ +lemma perfect.small_diam_splitting (ε_pos : 0 < ε) : ∃ C₀ C₁ : set α, + (perfect C₀ ∧ C₀.nonempty ∧ C₀ ⊆ C ∧ emetric.diam C₀ ≤ ε) ∧ + (perfect C₁ ∧ C₁.nonempty ∧ C₁ ⊆ C ∧ emetric.diam C₁ ≤ ε) ∧ disjoint C₀ C₁ := +begin + rcases hC.splitting hnonempty with ⟨D₀, D₁, ⟨perf0, non0, sub0⟩, ⟨perf1, non1, sub1⟩, hdisj⟩, + cases non0 with x₀ hx₀, + cases non1 with x₁ hx₁, + rcases perf0.small_diam_aux ε_pos hx₀ with ⟨perf0', non0', sub0', diam0⟩, + rcases perf1.small_diam_aux ε_pos hx₁ with ⟨perf1', non1', sub1', diam1⟩, + refine ⟨closure (emetric.ball x₀ (ε / 2) ∩ D₀), closure (emetric.ball x₁ (ε / 2) ∩ D₁), + ⟨perf0', non0', sub0'.trans sub0, diam0⟩, ⟨perf1', non1', sub1'.trans sub1, diam1⟩, _⟩, + apply disjoint.mono _ _ hdisj; assumption, +end + +open cantor_scheme + +/-- Any nonempty perfect set in a complete metric space admits a continuous injection +from the cantor space, `ℕ → bool`. -/ +theorem perfect.exists_nat_bool_injection [complete_space α] : + ∃ f : (ℕ → bool) → α, (range f) ⊆ C ∧ continuous f ∧ injective f := +begin + obtain ⟨u, -, upos', hu⟩ := exists_seq_strict_anti_tendsto' (zero_lt_one' ℝ≥0∞), + have upos := λ n, (upos' n).1, + let P := subtype (λ E : set α, perfect E ∧ E.nonempty), + choose C0 C1 h0 h1 hdisj using λ {C : set α} (hC : perfect C) (hnonempty : C.nonempty) + {ε : ℝ≥0∞} (hε : 0 < ε), hC.small_diam_splitting hnonempty hε, + let DP : list bool → P := λ l, + begin + induction l with a l ih, { exact ⟨C, ⟨hC, hnonempty⟩⟩ }, + cases a, + { use C0 ih.property.1 ih.property.2 (upos l.length.succ), + exact ⟨(h0 _ _ _).1, (h0 _ _ _).2.1⟩, }, + use C1 ih.property.1 ih.property.2 (upos l.length.succ), + exact ⟨(h1 _ _ _).1, (h1 _ _ _).2.1⟩, + end, + let D : list bool → set α := λ l, (DP l).val, + have hanti : closure_antitone D, + { refine antitone.closure_antitone _ (λ l, (DP l).property.1.closed), + intros l a, + cases a, + { exact (h0 _ _ _).2.2.1, }, + exact (h1 _ _ _).2.2.1, }, + have hdiam : vanishing_diam D, + { intro x, + apply tendsto_of_tendsto_of_tendsto_of_le_of_le' tendsto_const_nhds hu, + { simp }, + rw eventually_at_top, + refine ⟨1, λ m (hm : 1 ≤ m), _⟩, + rw nat.one_le_iff_ne_zero at hm, + rcases nat.exists_eq_succ_of_ne_zero hm with ⟨n, rfl⟩, + dsimp, + cases (x n), + { convert (h0 _ _ _).2.2.2, + rw pi_nat.res_length }, + convert (h1 _ _ _).2.2.2, + rw pi_nat.res_length, }, + have hdisj' : cantor_scheme.disjoint D, + { rintros l (a | a) (b | b) hab; try { contradiction }, + { exact hdisj _ _ _, }, + exact (hdisj _ _ _).symm, }, + have hdom : ∀ {x : ℕ → bool}, x ∈ (induced_map D).1 := λ x, + by simp [hanti.map_of_vanishing_diam hdiam (λ l, (DP l).property.2)], + refine ⟨λ x, (induced_map D).2 ⟨x, hdom⟩, _, _, _⟩, + { rintros y ⟨x, rfl⟩, + exact map_mem ⟨_, hdom⟩ 0, }, + { continuity, + exact hdiam.map_continuous, }, + intros x y hxy, + simpa only [← subtype.val_inj] using hdisj'.map_injective hxy, +end + +end cantor_inj From ce38d86c0b2d427ce208c3cee3159cb421d2b3c4 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Tue, 18 Apr 2023 05:13:33 +0000 Subject: [PATCH 0837/1291] chore(*): add mathlib4 synchronization comments (#18813) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.homology.homological_complex` * `algebra.homology.image_to_kernel` * `algebraic_topology.Moore_complex` * `algebraic_topology.topological_simplex` * `analysis.asymptotics.asymptotic_equivalent` * `analysis.box_integral.box.subbox_induction` * `analysis.convex.body` * `analysis.convex.exposed` * `analysis.normed.group.add_torsor` * `analysis.normed.ring.seminorm` * `analysis.normed_space.add_torsor` * `analysis.normed_space.pointwise` * `analysis.special_functions.polynomials` * `category_theory.sites.whiskering` * `category_theory.subobject.basic` * `category_theory.subobject.factor_thru` * `category_theory.subobject.lattice` * `category_theory.subobject.limits` * `category_theory.subobject.mono_over` * `category_theory.subobject.well_powered` * `data.num.prime` * `dynamics.circle.rotation_number.translation_number` * `geometry.manifold.local_invariant_properties` * `group_theory.free_abelian_group_finsupp` * `linear_algebra.free_module.finite.rank` * `model_theory.syntax` * `number_theory.lucas_lehmer` * `number_theory.pell_matiyasevic` * `order.category.DistLat` * `ring_theory.non_unital_subsemiring.basic` * `topology.algebra.nonarchimedean.bases` * `topology.metric_space.baire` * `topology.metric_space.closeds` --- src/algebra/homology/homological_complex.lean | 3 +++ src/algebra/homology/image_to_kernel.lean | 3 +++ src/algebraic_topology/Moore_complex.lean | 3 +++ src/algebraic_topology/topological_simplex.lean | 3 +++ src/analysis/asymptotics/asymptotic_equivalent.lean | 3 +++ src/analysis/box_integral/box/subbox_induction.lean | 3 +++ src/analysis/convex/body.lean | 3 +++ src/analysis/convex/exposed.lean | 3 +++ src/analysis/normed/group/add_torsor.lean | 3 +++ src/analysis/normed/ring/seminorm.lean | 3 +++ src/analysis/normed_space/add_torsor.lean | 3 +++ src/analysis/normed_space/pointwise.lean | 3 +++ src/analysis/special_functions/polynomials.lean | 3 +++ src/category_theory/sites/whiskering.lean | 3 +++ src/category_theory/subobject/basic.lean | 3 +++ src/category_theory/subobject/factor_thru.lean | 3 +++ src/category_theory/subobject/lattice.lean | 3 +++ src/category_theory/subobject/limits.lean | 3 +++ src/category_theory/subobject/mono_over.lean | 3 +++ src/category_theory/subobject/well_powered.lean | 3 +++ src/data/num/prime.lean | 3 +++ src/dynamics/circle/rotation_number/translation_number.lean | 3 +++ src/geometry/manifold/local_invariant_properties.lean | 3 +++ src/group_theory/free_abelian_group_finsupp.lean | 3 +++ src/linear_algebra/free_module/finite/rank.lean | 3 +++ src/model_theory/syntax.lean | 3 +++ src/number_theory/lucas_lehmer.lean | 3 +++ src/number_theory/pell_matiyasevic.lean | 3 +++ src/order/category/DistLat.lean | 3 +++ src/ring_theory/non_unital_subsemiring/basic.lean | 3 +++ src/topology/algebra/nonarchimedean/bases.lean | 3 +++ src/topology/metric_space/baire.lean | 3 +++ src/topology/metric_space/closeds.lean | 3 +++ 33 files changed, 99 insertions(+) diff --git a/src/algebra/homology/homological_complex.lean b/src/algebra/homology/homological_complex.lean index c948c649e5010..7bbcf28f61770 100644 --- a/src/algebra/homology/homological_complex.lean +++ b/src/algebra/homology/homological_complex.lean @@ -10,6 +10,9 @@ import category_theory.graded_object /-! # Homological complexes. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A `homological_complex V c` with a "shape" controlled by `c : complex_shape ι` has chain groups `X i` (objects in `V`) indexed by `i : ι`, and a differential `d i j` whenever `c.rel i j`. diff --git a/src/algebra/homology/image_to_kernel.lean b/src/algebra/homology/image_to_kernel.lean index e890a3d2db301..5b100a5aec890 100644 --- a/src/algebra/homology/image_to_kernel.lean +++ b/src/algebra/homology/image_to_kernel.lean @@ -8,6 +8,9 @@ import category_theory.subobject.limits /-! # Image-to-kernel comparison maps +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Whenever `f : A ⟶ B` and `g : B ⟶ C` satisfy `w : f ≫ g = 0`, we have `image_le_kernel f g w : image_subobject f ≤ kernel_subobject g` (assuming the appropriate images and kernels exist). diff --git a/src/algebraic_topology/Moore_complex.lean b/src/algebraic_topology/Moore_complex.lean index 79808b2f721a1..5c835034f90df 100644 --- a/src/algebraic_topology/Moore_complex.lean +++ b/src/algebraic_topology/Moore_complex.lean @@ -10,6 +10,9 @@ import category_theory.abelian.basic /-! ## Moore complex +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We construct the normalized Moore complex, as a functor `simplicial_object C ⥤ chain_complex C ℕ`, for any abelian category `C`. diff --git a/src/algebraic_topology/topological_simplex.lean b/src/algebraic_topology/topological_simplex.lean index 87a07cc883be3..552095a9fa087 100644 --- a/src/algebraic_topology/topological_simplex.lean +++ b/src/algebraic_topology/topological_simplex.lean @@ -10,6 +10,9 @@ import topology.instances.nnreal /-! # Topological simplices +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define the natural functor from `simplex_category` to `Top` sending `[n]` to the topological `n`-simplex. This is used to define `Top.to_sSet` in `algebraic_topology.simpliciaL_set`. diff --git a/src/analysis/asymptotics/asymptotic_equivalent.lean b/src/analysis/asymptotics/asymptotic_equivalent.lean index 5d50ec2ce1c0b..3eb177685acb9 100644 --- a/src/analysis/asymptotics/asymptotic_equivalent.lean +++ b/src/analysis/asymptotics/asymptotic_equivalent.lean @@ -9,6 +9,9 @@ import analysis.normed.order.basic /-! # Asymptotic equivalence +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we define the relation `is_equivalent l u v`, which means that `u-v` is little o of `v` along the filter `l`. diff --git a/src/analysis/box_integral/box/subbox_induction.lean b/src/analysis/box_integral/box/subbox_induction.lean index 1f25231c4836d..c3aa323169157 100644 --- a/src/analysis/box_integral/box/subbox_induction.lean +++ b/src/analysis/box_integral/box/subbox_induction.lean @@ -9,6 +9,9 @@ import analysis.specific_limits.basic /-! # Induction on subboxes +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove the following induction principle for `box_integral.box`, see `box_integral.box.subbox_induction_on`. Let `p` be a predicate on `box_integral.box ι`, let `I` be a box. Suppose that the following two properties hold true. diff --git a/src/analysis/convex/body.lean b/src/analysis/convex/body.lean index 8a8698e4986b1..bba94350b7b54 100644 --- a/src/analysis/convex/body.lean +++ b/src/analysis/convex/body.lean @@ -10,6 +10,9 @@ import topology.metric_space.hausdorff_distance /-! # Convex bodies +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains the definition of the type `convex_body V` consisting of convex, compact, nonempty subsets of a real topological vector space `V`. diff --git a/src/analysis/convex/exposed.lean b/src/analysis/convex/exposed.lean index 65bd344a1ae25..c8e7f15e4f45b 100644 --- a/src/analysis/convex/exposed.lean +++ b/src/analysis/convex/exposed.lean @@ -11,6 +11,9 @@ import topology.order.basic /-! # Exposed sets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines exposed sets and exposed points for sets in a real vector space. An exposed subset of `A` is a subset of `A` that is the set of all maximal points of a functional diff --git a/src/analysis/normed/group/add_torsor.lean b/src/analysis/normed/group/add_torsor.lean index 4a526f6fa62a1..169615af036ed 100644 --- a/src/analysis/normed/group/add_torsor.lean +++ b/src/analysis/normed/group/add_torsor.lean @@ -10,6 +10,9 @@ import linear_algebra.affine_space.midpoint /-! # Torsors of additive normed group actions. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines torsors of additive normed group actions, with a metric space structure. The motivating case is Euclidean affine spaces. diff --git a/src/analysis/normed/ring/seminorm.lean b/src/analysis/normed/ring/seminorm.lean index c42154bb990a3..21c6fbd9261b7 100644 --- a/src/analysis/normed/ring/seminorm.lean +++ b/src/analysis/normed/ring/seminorm.lean @@ -8,6 +8,9 @@ import analysis.normed.field.basic /-! # Seminorms and norms on rings +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines seminorms and norms on rings. These definitions are useful when one needs to consider multiple (semi)norms on a given ring. diff --git a/src/analysis/normed_space/add_torsor.lean b/src/analysis/normed_space/add_torsor.lean index 72d42650c7308..001c01845a245 100644 --- a/src/analysis/normed_space/add_torsor.lean +++ b/src/analysis/normed_space/add_torsor.lean @@ -12,6 +12,9 @@ import topology.instances.real_vector_space /-! # Torsors of normed space actions. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains lemmas about normed additive torsors over normed spaces. -/ diff --git a/src/analysis/normed_space/pointwise.lean b/src/analysis/normed_space/pointwise.lean index a236e8efe884e..65e8d118b8b47 100644 --- a/src/analysis/normed_space/pointwise.lean +++ b/src/analysis/normed_space/pointwise.lean @@ -10,6 +10,9 @@ import analysis.normed_space.basic /-! # Properties of pointwise scalar multiplication of sets in normed spaces. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We explore the relationships between scalar multiplication of sets in vector spaces, and the norm. Notably, we express arbitrary balls as rescaling of other balls, and we show that the multiplication of bounded sets remain bounded. diff --git a/src/analysis/special_functions/polynomials.lean b/src/analysis/special_functions/polynomials.lean index 41b3c39a452b1..779b88c4163d8 100644 --- a/src/analysis/special_functions/polynomials.lean +++ b/src/analysis/special_functions/polynomials.lean @@ -10,6 +10,9 @@ import data.polynomial.ring_division /-! # Limits related to polynomial and rational functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves basic facts about limits of polynomial and rationals functions. The main result is `eval_is_equivalent_at_top_eval_lead`, which states that for any polynomial `P` of degree `n` with leading coefficient `a`, the corresponding diff --git a/src/category_theory/sites/whiskering.lean b/src/category_theory/sites/whiskering.lean index 595368886b0d1..c437b9d5be03c 100644 --- a/src/category_theory/sites/whiskering.lean +++ b/src/category_theory/sites/whiskering.lean @@ -7,6 +7,9 @@ Authors: Adam Topaz import category_theory.sites.sheaf /-! +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we construct the functor `Sheaf J A ⥤ Sheaf J B` between sheaf categories obtained by composition with a functor `F : A ⥤ B`. diff --git a/src/category_theory/subobject/basic.lean b/src/category_theory/subobject/basic.lean index be79f761748f6..c32132d2d2711 100644 --- a/src/category_theory/subobject/basic.lean +++ b/src/category_theory/subobject/basic.lean @@ -12,6 +12,9 @@ import tactic.elementwise /-! # Subobjects +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define `subobject X` as the quotient (by isomorphisms) of `mono_over X := {f : over X // mono f.hom}`. diff --git a/src/category_theory/subobject/factor_thru.lean b/src/category_theory/subobject/factor_thru.lean index c5f450340f18c..6d1fe817d6a5c 100644 --- a/src/category_theory/subobject/factor_thru.lean +++ b/src/category_theory/subobject/factor_thru.lean @@ -9,6 +9,9 @@ import category_theory.preadditive.basic /-! # Factoring through subobjects +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The predicate `h : P.factors f`, for `P : subobject Y` and `f : X ⟶ Y` asserts the existence of some `P.factor_thru f : X ⟶ (P : C)` making the obvious diagram commute. diff --git a/src/category_theory/subobject/lattice.lean b/src/category_theory/subobject/lattice.lean index 126de21d00991..319f87ae235ef 100644 --- a/src/category_theory/subobject/lattice.lean +++ b/src/category_theory/subobject/lattice.lean @@ -9,6 +9,9 @@ import category_theory.subobject.well_powered /-! # The lattice of subobjects +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We provide the `semilattice_inf` with `order_top (subobject X)` instance when `[has_pullback C]`, and the `semilattice_sup (subobject X)` instance when `[has_images C] [has_binary_coproducts C]`. -/ diff --git a/src/category_theory/subobject/limits.lean b/src/category_theory/subobject/limits.lean index d449eaaf73ba6..08e559e0c2913 100644 --- a/src/category_theory/subobject/limits.lean +++ b/src/category_theory/subobject/limits.lean @@ -8,6 +8,9 @@ import category_theory.subobject.lattice /-! # Specific subobjects +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define `equalizer_subobject`, `kernel_subobject` and `image_subobject`, which are the subobjects represented by the equalizer, kernel and image of (a pair of) morphism(s) and provide conditions for `P.factors f`, where `P` is one of these special subobjects. diff --git a/src/category_theory/subobject/mono_over.lean b/src/category_theory/subobject/mono_over.lean index 87410369a593d..619c45c78ae36 100644 --- a/src/category_theory/subobject/mono_over.lean +++ b/src/category_theory/subobject/mono_over.lean @@ -10,6 +10,9 @@ import category_theory.adjunction.reflective /-! # Monomorphisms over a fixed object +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + As preparation for defining `subobject X`, we set up the theory for `mono_over X := {f : over X // mono f.hom}`. diff --git a/src/category_theory/subobject/well_powered.lean b/src/category_theory/subobject/well_powered.lean index 9652b24e8a0c8..3336ecda1e177 100644 --- a/src/category_theory/subobject/well_powered.lean +++ b/src/category_theory/subobject/well_powered.lean @@ -9,6 +9,9 @@ import category_theory.essentially_small /-! # Well-powered categories +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A category `(C : Type u) [category.{v} C]` is `[well_powered C]` if for every `X : C`, we have `small.{v} (subobject X)`. diff --git a/src/data/num/prime.lean b/src/data/num/prime.lean index 08ee6312966d1..f996feac90da0 100644 --- a/src/data/num/prime.lean +++ b/src/data/num/prime.lean @@ -10,6 +10,9 @@ import tactic.ring /-! # Primality for binary natural numbers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines versions of `nat.min_fac` and `nat.prime` for `num` and `pos_num`. As with other `num` definitions, they are not intended for general use (`nat` should be used instead of `num` in most cases) but they can be used in contexts where kernel computation is required, such as proofs diff --git a/src/dynamics/circle/rotation_number/translation_number.lean b/src/dynamics/circle/rotation_number/translation_number.lean index 0851e653cab39..4c340988b2134 100644 --- a/src/dynamics/circle/rotation_number/translation_number.lean +++ b/src/dynamics/circle/rotation_number/translation_number.lean @@ -12,6 +12,9 @@ import topology.algebra.order.monotone_continuity /-! # Translation number of a monotone real map that commutes with `x ↦ x + 1` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Let `f : ℝ → ℝ` be a monotone map such that `f (x + 1) = f x + 1` for all `x`. Then the limit $$ \tau(f)=\lim_{n\to\infty}{f^n(x)-x}{n} diff --git a/src/geometry/manifold/local_invariant_properties.lean b/src/geometry/manifold/local_invariant_properties.lean index 6136f52a119f1..8e07981ff9550 100644 --- a/src/geometry/manifold/local_invariant_properties.lean +++ b/src/geometry/manifold/local_invariant_properties.lean @@ -8,6 +8,9 @@ import geometry.manifold.charted_space /-! # Local properties invariant under a groupoid +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We study properties of a triple `(g, s, x)` where `g` is a function between two spaces `H` and `H'`, `s` is a subset of `H` and `x` is a point of `H`. Our goal is to register how such a property should behave to make sense in charted spaces modelled on `H` and `H'`. diff --git a/src/group_theory/free_abelian_group_finsupp.lean b/src/group_theory/free_abelian_group_finsupp.lean index 4eaefccabea1f..10e05a6d2520b 100644 --- a/src/group_theory/free_abelian_group_finsupp.lean +++ b/src/group_theory/free_abelian_group_finsupp.lean @@ -14,6 +14,9 @@ import linear_algebra.dimension /-! # Isomorphism between `free_abelian_group X` and `X →₀ ℤ` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we construct the canonical isomorphism between `free_abelian_group X` and `X →₀ ℤ`. We use this to transport the notion of `support` from `finsupp` to `free_abelian_group`. diff --git a/src/linear_algebra/free_module/finite/rank.lean b/src/linear_algebra/free_module/finite/rank.lean index 1b2ace77c9a8b..0de26d8d1ba11 100644 --- a/src/linear_algebra/free_module/finite/rank.lean +++ b/src/linear_algebra/free_module/finite/rank.lean @@ -12,6 +12,9 @@ import linear_algebra.free_module.finite.basic # Rank of finite free modules +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This is a basic API for the rank of finite free modules. -/ diff --git a/src/model_theory/syntax.lean b/src/model_theory/syntax.lean index 80477f6c59bbf..288d455200966 100644 --- a/src/model_theory/syntax.lean +++ b/src/model_theory/syntax.lean @@ -10,6 +10,9 @@ import model_theory.language_map /-! # Basics on First-Order Syntax + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. This file defines first-order terms, formulas, sentences, and theories in a style inspired by the [Flypitch project](https://flypitch.github.io/). diff --git a/src/number_theory/lucas_lehmer.lean b/src/number_theory/lucas_lehmer.lean index d0d8613cb17d3..f05fc1a16b458 100644 --- a/src/number_theory/lucas_lehmer.lean +++ b/src/number_theory/lucas_lehmer.lean @@ -14,6 +14,9 @@ import tactic.ring_exp /-! # The Lucas-Lehmer test for Mersenne primes. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define `lucas_lehmer_residue : Π p : ℕ, zmod (2^p - 1)`, and prove `lucas_lehmer_residue p = 0 → prime (mersenne p)`. diff --git a/src/number_theory/pell_matiyasevic.lean b/src/number_theory/pell_matiyasevic.lean index 76830c1fc9049..0884c8773a3b1 100644 --- a/src/number_theory/pell_matiyasevic.lean +++ b/src/number_theory/pell_matiyasevic.lean @@ -11,6 +11,9 @@ import number_theory.zsqrtd.basic /-! # Pell's equation and Matiyasevic's theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file solves Pell's equation, i.e. integer solutions to `x ^ 2 - d * y ^ 2 = 1` *in the special case that `d = a ^ 2 - 1`*. This is then applied to prove Matiyasevic's theorem that the power diff --git a/src/order/category/DistLat.lean b/src/order/category/DistLat.lean index c351ee58c40db..c4633bbbe1434 100644 --- a/src/order/category/DistLat.lean +++ b/src/order/category/DistLat.lean @@ -8,6 +8,9 @@ import order.category.Lat /-! # The category of distributive lattices +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines `DistLat`, the category of distributive lattices. Note that [`DistLat`](https://ncatlab.org/nlab/show/DistLat) in the literature doesn't always diff --git a/src/ring_theory/non_unital_subsemiring/basic.lean b/src/ring_theory/non_unital_subsemiring/basic.lean index 5d85dd07d2a9e..2368f132629a9 100644 --- a/src/ring_theory/non_unital_subsemiring/basic.lean +++ b/src/ring_theory/non_unital_subsemiring/basic.lean @@ -14,6 +14,9 @@ import group_theory.subsemigroup.centralizer /-! # Bundled non-unital subsemirings +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define bundled non-unital subsemirings and some standard constructions: `complete_lattice` structure, `subtype` and `inclusion` ring homomorphisms, non-unital subsemiring `map`, `comap` and range (`srange`) of a `non_unital_ring_hom` etc. diff --git a/src/topology/algebra/nonarchimedean/bases.lean b/src/topology/algebra/nonarchimedean/bases.lean index d1a63ae9a582b..29401a6db5294 100644 --- a/src/topology/algebra/nonarchimedean/bases.lean +++ b/src/topology/algebra/nonarchimedean/bases.lean @@ -11,6 +11,9 @@ import algebra.module.submodule.pointwise /-! # Neighborhood bases for non-archimedean rings and modules +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This files contains special families of filter bases on rings and modules that give rise to non-archimedean topologies. diff --git a/src/topology/metric_space/baire.lean b/src/topology/metric_space/baire.lean index 7de9a219c1aec..348298fedc1b2 100644 --- a/src/topology/metric_space/baire.lean +++ b/src/topology/metric_space/baire.lean @@ -11,6 +11,9 @@ import topology.sets.compacts /-! # Baire theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In a complete metric space, a countable intersection of dense open subsets is dense. The good concept underlying the theorem is that of a Gδ set, i.e., a countable intersection diff --git a/src/topology/metric_space/closeds.lean b/src/topology/metric_space/closeds.lean index 12e5998f38e63..fa1cdd15b9b8d 100644 --- a/src/topology/metric_space/closeds.lean +++ b/src/topology/metric_space/closeds.lean @@ -10,6 +10,9 @@ import topology.sets.compacts /-! # Closed subsets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the metric and emetric space structure on the types of closed subsets and nonempty compact subsets of a metric or emetric space. From 6c263e4bfc2e6714de30f22178b4d0ca4d149a76 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Tue, 18 Apr 2023 07:20:54 +0000 Subject: [PATCH 0838/1291] chore(linear_algebra/matrix/basis): reduce imports (#18823) This is a slight simplification on the current maximal path for the port. Co-authored-by: Scott Morrison --- src/linear_algebra/matrix/basis.lean | 3 +-- 1 file changed, 1 insertion(+), 2 deletions(-) diff --git a/src/linear_algebra/matrix/basis.lean b/src/linear_algebra/matrix/basis.lean index bd9c8e5356b06..2872c03cf147b 100644 --- a/src/linear_algebra/matrix/basis.lean +++ b/src/linear_algebra/matrix/basis.lean @@ -3,7 +3,6 @@ Copyright (c) 2019 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Patrick Massot, Casper Putz, Anne Baanen -/ -import linear_algebra.matrix.nonsingular_inverse import linear_algebra.matrix.reindex import linear_algebra.matrix.to_lin @@ -239,7 +238,7 @@ by rw [basis.to_matrix_mul_to_matrix, basis.to_matrix_self] /-- A matrix whose columns form a basis `b'`, expressed w.r.t. a basis `b`, is invertible. -/ def basis.invertible_to_matrix [decidable_eq ι] [fintype ι] (b b' : basis ι R₂ M₂) : invertible (b.to_matrix b') := -matrix.invertible_of_left_inverse _ _ (basis.to_matrix_mul_to_matrix_flip _ _) +⟨b'.to_matrix b, basis.to_matrix_mul_to_matrix_flip _ _, basis.to_matrix_mul_to_matrix_flip _ _⟩ @[simp] lemma basis.to_matrix_reindex From cd8fafa2fac98e1a67097e8a91ad9901cfde48af Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Tue, 18 Apr 2023 07:20:56 +0000 Subject: [PATCH 0839/1291] chore(data/complex/module): split out orientation to a separate file (#18824) This removes the imports - linear_algebra.matrix.determinant - linear_algebra.matrix.mv_polynomial - linear_algebra.matrix.adjugate - linear_algebra.matrix.nonsingular_inverse - linear_algebra.matrix.basis - linear_algebra.determinant - linear_algebra.orientation from `data.complex.module`, which aren't otherwise needed here. Co-authored-by: Scott Morrison --- src/analysis/inner_product_space/two_dim.lean | 1 + src/data/complex/module.lean | 4 ---- src/data/complex/orientation.lean | 21 +++++++++++++++++++ 3 files changed, 22 insertions(+), 4 deletions(-) create mode 100644 src/data/complex/orientation.lean diff --git a/src/analysis/inner_product_space/two_dim.lean b/src/analysis/inner_product_space/two_dim.lean index a55a4456334a1..62fa50cecc64e 100644 --- a/src/analysis/inner_product_space/two_dim.lean +++ b/src/analysis/inner_product_space/two_dim.lean @@ -5,6 +5,7 @@ Authors: Heather Macbeth -/ import analysis.inner_product_space.dual import analysis.inner_product_space.orientation +import data.complex.orientation import tactic.linear_combination /-! diff --git a/src/data/complex/module.lean b/src/data/complex/module.lean index d3a67b8891069..b17fd2e2a7087 100644 --- a/src/data/complex/module.lean +++ b/src/data/complex/module.lean @@ -3,7 +3,6 @@ Copyright (c) 2020 Alexander Bentkamp, Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Alexander Bentkamp, Sébastien Gouëzel, Eric Wieser -/ -import linear_algebra.orientation import algebra.order.smul import data.complex.basic import data.fin.vec_notation @@ -167,9 +166,6 @@ by simp [← finrank_eq_rank, finrank_real_complex, bit0] circle. -/ lemma finrank_real_complex_fact : fact (finrank ℝ ℂ = 2) := ⟨finrank_real_complex⟩ -/-- The standard orientation on `ℂ`. -/ -protected noncomputable def orientation : orientation ℝ ℂ (fin 2) := complex.basis_one_I.orientation - end complex /- Register as an instance (with low priority) the fact that a complex vector space is also a real diff --git a/src/data/complex/orientation.lean b/src/data/complex/orientation.lean new file mode 100644 index 0000000000000..f2ba3d19f2737 --- /dev/null +++ b/src/data/complex/orientation.lean @@ -0,0 +1,21 @@ +/- +Copyright (c) 2021 Heather Macbeth. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Heather Macbeth +-/ +import data.complex.module +import linear_algebra.orientation + +/-! +# The standard orientation on `ℂ`. + +This had previously been in `linear_algebra.orientation`, +but keeping it separate results in a significant import reduction. +-/ + +namespace complex + +/-- The standard orientation on `ℂ`. -/ +protected noncomputable def orientation : orientation ℝ ℂ (fin 2) := complex.basis_one_I.orientation + +end complex From c6275ef4bef8a44472109311361a0eacae160e1e Mon Sep 17 00:00:00 2001 From: lukemantle Date: Wed, 19 Apr 2023 08:40:29 +0000 Subject: [PATCH 0840/1291] feat(ring_theory/polynomial): define the probabilist's Hermite polynomials (#18739) Define the Hermite polynomials recursively, and show this is equivalent to the result of iterating the operation `x - d/dx` on the constant polynomial `1`. Future PRs will include several equivalent characterizations: - Recursive and explicit expressions for the coefficients - Definition based on the nth derivative of the Gaussian function Co-authored-by: Jake Levinson Co-authored-by: lukemantle <62187700+lukemantle@users.noreply.github.com> Co-authored-by: Eric Wieser --- src/ring_theory/polynomial/hermite.lean | 53 +++++++++++++++++++++++++ 1 file changed, 53 insertions(+) create mode 100644 src/ring_theory/polynomial/hermite.lean diff --git a/src/ring_theory/polynomial/hermite.lean b/src/ring_theory/polynomial/hermite.lean new file mode 100644 index 0000000000000..a1a5d7f677c91 --- /dev/null +++ b/src/ring_theory/polynomial/hermite.lean @@ -0,0 +1,53 @@ +/- +Copyright (c) 2023 Luke Mantle. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Luke Mantle +-/ + +import data.polynomial.derivative + +/-! +# Hermite polynomials + +This file defines `polynomial.hermite n`, the nth probabilist's Hermite polynomial. + +## Main definitions + +* `polynomial.hermite n`: the `n`th probabilist's Hermite polynomial, + defined recursively as a `polynomial ℤ` + +## References + +* [Hermite Polynomials](https://en.wikipedia.org/wiki/Hermite_polynomials) + +-/ + +noncomputable theory +open polynomial + +namespace polynomial + +/-- the nth probabilist's Hermite polynomial -/ +noncomputable def hermite : ℕ → polynomial ℤ +| 0 := 1 +| (n+1) := X * (hermite n) - (hermite n).derivative + +@[simp] lemma hermite_succ (n : ℕ) : hermite (n+1) = X * (hermite n) - (hermite n).derivative := +by rw hermite + +lemma hermite_eq_iterate (n : ℕ) : hermite n = ((λ p, X*p - p.derivative)^[n] 1) := +begin + induction n with n ih, + { refl }, + { rw [function.iterate_succ_apply', ← ih, hermite_succ] } +end + +@[simp] lemma hermite_zero : hermite 0 = C 1 := rfl + +@[simp] lemma hermite_one : hermite 1 = X := +begin + rw [hermite_succ, hermite_zero], + simp only [map_one, mul_one, derivative_one, sub_zero] +end + +end polynomial From 8eb9c42d4d34c77f6ee84ea766ae4070233a973c Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Wed, 19 Apr 2023 14:57:28 +0000 Subject: [PATCH 0841/1291] chore(*): add mathlib4 synchronization comments (#18832) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.homology.exact` * `algebra.homology.flip` * `algebra.homology.functor` * `algebra.homology.homology` * `algebra.homology.single` * `category_theory.monad.types` * `category_theory.subobject.types` * `computability.tm_computable` * `control.bitraversable.lemmas` * `data.qpf.univariate.basic` * `data.real.cardinality` * `model_theory.encoding` * `ring_theory.localization.ideal` * `topology.algebra.valuation` * `topology.metric_space.cantor_scheme` * `topology.urysohns_lemma` --- src/algebra/homology/exact.lean | 3 +++ src/algebra/homology/flip.lean | 3 +++ src/algebra/homology/functor.lean | 3 +++ src/algebra/homology/homology.lean | 3 +++ src/algebra/homology/single.lean | 3 +++ src/category_theory/monad/types.lean | 3 +++ src/category_theory/subobject/types.lean | 3 +++ src/computability/tm_computable.lean | 3 +++ src/control/bitraversable/lemmas.lean | 3 +++ src/data/qpf/univariate/basic.lean | 3 +++ src/data/real/cardinality.lean | 3 +++ src/model_theory/encoding.lean | 3 +++ src/ring_theory/localization/ideal.lean | 3 +++ src/topology/algebra/valuation.lean | 3 +++ src/topology/metric_space/cantor_scheme.lean | 3 +++ src/topology/urysohns_lemma.lean | 3 +++ 16 files changed, 48 insertions(+) diff --git a/src/algebra/homology/exact.lean b/src/algebra/homology/exact.lean index 4f29c32770089..a38897e0d7ada 100644 --- a/src/algebra/homology/exact.lean +++ b/src/algebra/homology/exact.lean @@ -8,6 +8,9 @@ import algebra.homology.image_to_kernel /-! # Exact sequences +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In a category with zero morphisms, images, and equalizers we say that `f : A ⟶ B` and `g : B ⟶ C` are exact if `f ≫ g = 0` and the natural map `image f ⟶ kernel g` is an epimorphism. diff --git a/src/algebra/homology/flip.lean b/src/algebra/homology/flip.lean index 8963be0280c5d..310b4ac89443a 100644 --- a/src/algebra/homology/flip.lean +++ b/src/algebra/homology/flip.lean @@ -8,6 +8,9 @@ import algebra.homology.homological_complex /-! # Flip a complex of complexes +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + For now we don't have double complexes as a distinct thing, but we can model them as complexes of complexes. diff --git a/src/algebra/homology/functor.lean b/src/algebra/homology/functor.lean index 87f4bc213bacc..cee1f32ec5ddd 100644 --- a/src/algebra/homology/functor.lean +++ b/src/algebra/homology/functor.lean @@ -8,6 +8,9 @@ import algebra.homology.homological_complex /-! # Complexes in functor categories +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We can view a complex valued in a functor category `T ⥤ V` as a functor from `T` to complexes valued in `V`. diff --git a/src/algebra/homology/homology.lean b/src/algebra/homology/homology.lean index 30166c8f1f5ef..348197a115a68 100644 --- a/src/algebra/homology/homology.lean +++ b/src/algebra/homology/homology.lean @@ -10,6 +10,9 @@ import category_theory.graded_object /-! # The homology of a complex +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given `C : homological_complex V c`, we have `C.cycles i` and `C.boundaries i`, both defined as subobjects of `C.X i`. diff --git a/src/algebra/homology/single.lean b/src/algebra/homology/single.lean index dee21c6c21614..ca25cc2f960e9 100644 --- a/src/algebra/homology/single.lean +++ b/src/algebra/homology/single.lean @@ -8,6 +8,9 @@ import algebra.homology.homology /-! # Chain complexes supported in a single degree +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define `single V j c : V ⥤ homological_complex V c`, which constructs complexes in `V` of shape `c`, supported in degree `j`. diff --git a/src/category_theory/monad/types.lean b/src/category_theory/monad/types.lean index 962cb2e07c726..f231851f08ff5 100644 --- a/src/category_theory/monad/types.lean +++ b/src/category_theory/monad/types.lean @@ -12,6 +12,9 @@ import category_theory.types # Convert from `monad` (i.e. Lean's `Type`-based monads) to `category_theory.monad` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This allows us to use these monads in category theory. -/ diff --git a/src/category_theory/subobject/types.lean b/src/category_theory/subobject/types.lean index a36e60ce02444..f53ee33ba1274 100644 --- a/src/category_theory/subobject/types.lean +++ b/src/category_theory/subobject/types.lean @@ -9,6 +9,9 @@ import category_theory.types /-! # `Type u` is well-powered +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + By building a categorical equivalence `mono_over α ≌ set α` for any `α : Type u`, we deduce that `subobject α ≃o set α` and that `Type u` is well-powered. diff --git a/src/computability/tm_computable.lean b/src/computability/tm_computable.lean index 8a62da32aac3d..45b5ada4d0133 100644 --- a/src/computability/tm_computable.lean +++ b/src/computability/tm_computable.lean @@ -12,6 +12,9 @@ import data.polynomial.eval /-! # Computable functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains the definition of a Turing machine with some finiteness conditions (bundling the definition of TM2 in turing_machine.lean), a definition of when a TM gives a certain output (in a certain time), and the definition of computability (in polytime or any time function) diff --git a/src/control/bitraversable/lemmas.lean b/src/control/bitraversable/lemmas.lean index d668f808be5b0..b2f76e42b630a 100644 --- a/src/control/bitraversable/lemmas.lean +++ b/src/control/bitraversable/lemmas.lean @@ -8,6 +8,9 @@ import control.bitraversable.basic /-! # Bitraversable Lemmas +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions * tfst - traverse on first functor argument * tsnd - traverse on second functor argument diff --git a/src/data/qpf/univariate/basic.lean b/src/data/qpf/univariate/basic.lean index 94c474a2a2055..0d83b3146acdb 100644 --- a/src/data/qpf/univariate/basic.lean +++ b/src/data/qpf/univariate/basic.lean @@ -9,6 +9,9 @@ import data.pfunctor.univariate.M # Quotients of Polynomial Functors +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We assume the following: `P` : a polynomial functor diff --git a/src/data/real/cardinality.lean b/src/data/real/cardinality.lean index 74c5933f30a7f..189427b8ebc41 100644 --- a/src/data/real/cardinality.lean +++ b/src/data/real/cardinality.lean @@ -11,6 +11,9 @@ import set_theory.cardinal.continuum /-! # The cardinality of the reals +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file shows that the real numbers have cardinality continuum, i.e. `#ℝ = 𝔠`. We show that `#ℝ ≤ 𝔠` by noting that every real number is determined by a Cauchy-sequence of the diff --git a/src/model_theory/encoding.lean b/src/model_theory/encoding.lean index b23894d20b87e..68e4c9e97bf9f 100644 --- a/src/model_theory/encoding.lean +++ b/src/model_theory/encoding.lean @@ -11,6 +11,9 @@ import set_theory.cardinal.ordinal /-! # Encodings and Cardinality of First-Order Syntax +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main Definitions * `first_order.language.term.encoding` encodes terms as lists. * `first_order.language.bounded_formula.encoding` encodes bounded formulas as lists. diff --git a/src/ring_theory/localization/ideal.lean b/src/ring_theory/localization/ideal.lean index c67487279fc2d..93a871f3285fd 100644 --- a/src/ring_theory/localization/ideal.lean +++ b/src/ring_theory/localization/ideal.lean @@ -9,6 +9,9 @@ import ring_theory.localization.basic /-! # Ideals in localizations of commutative rings +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Implementation notes See `src/ring_theory/localization/basic.lean` for a design overview. diff --git a/src/topology/algebra/valuation.lean b/src/topology/algebra/valuation.lean index 737093c2420b4..a89243493d4c3 100644 --- a/src/topology/algebra/valuation.lean +++ b/src/topology/algebra/valuation.lean @@ -11,6 +11,9 @@ import ring_theory.valuation.basic /-! # The topology on a valued ring +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we define the non archimedean topology induced by a valuation on a ring. The main definition is a `valued` type class which equips a ring with a valuation taking values in a group with zero. Other instances are then deduced from this. diff --git a/src/topology/metric_space/cantor_scheme.lean b/src/topology/metric_space/cantor_scheme.lean index 276a911ae5514..eea5317cc3187 100644 --- a/src/topology/metric_space/cantor_scheme.lean +++ b/src/topology/metric_space/cantor_scheme.lean @@ -8,6 +8,9 @@ import topology.metric_space.pi_nat /-! # (Topological) Schemes and their induced maps +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In topology, and especially descriptive set theory, one often constructs functions `(ℕ → β) → α`, where α is some topological space and β is a discrete space, as an appropriate limit of some map `list β → set α`. We call the latter type of map a "`β`-scheme on `α`". diff --git a/src/topology/urysohns_lemma.lean b/src/topology/urysohns_lemma.lean index e9236a07a162f..86952bc44ec7b 100644 --- a/src/topology/urysohns_lemma.lean +++ b/src/topology/urysohns_lemma.lean @@ -10,6 +10,9 @@ import topology.continuous_function.basic /-! # Urysohn's lemma +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove Urysohn's lemma `exists_continuous_zero_one_of_closed`: for any two disjoint closed sets `s` and `t` in a normal topological space `X` there exists a continuous function `f : X → ℝ` such that From 17219820a8aa8abe85adf5dfde19af1dd1bd8ae7 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 19 Apr 2023 16:38:57 +0000 Subject: [PATCH 0842/1291] =?UTF-8?q?feat(data/matrix/rank):=20rank=20of?= =?UTF-8?q?=20`A=E1=B5=80=20=E2=AC=9D=20A`=20and=20`A=E1=B4=B4=20=E2=AC=9D?= =?UTF-8?q?=20A`=20(#18818)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Since these results imply it trivially, this also includes lemmas about the rank of `Aᵀ` and `Aᴴ`. However, these lemmas are not stated very generally, and are surely true in wider cases than the ones proven here. --- src/data/matrix/rank.lean | 120 ++++++++++++++++++++++++++++++++++++-- 1 file changed, 116 insertions(+), 4 deletions(-) diff --git a/src/data/matrix/rank.lean b/src/data/matrix/rank.lean index 1e491bb9adbfe..2caf7ce437817 100644 --- a/src/data/matrix/rank.lean +++ b/src/data/matrix/rank.lean @@ -1,11 +1,14 @@ /- Copyright (c) 2021 Johan Commelin. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Johan Commelin +Authors: Johan Commelin, Eric Wieer -/ import linear_algebra.free_module.finite.rank import linear_algebra.matrix.to_lin +import linear_algebra.finite_dimensional +import linear_algebra.matrix.dot_product +import data.complex.module /-! # Rank of matrices @@ -19,7 +22,9 @@ This definition does not depend on the choice of basis, see `matrix.rank_eq_finr ## TODO -* Show that `matrix.rank` is equal to the row-rank, and that `rank Aᵀ = rank A`. +* Do a better job of generalizing over `ℚ`, `ℝ`, and `ℂ` in `matrix.rank_transpose` and + `matrix.rank_conj_transpose`. See + [this Zulip thread](https://leanprover.zulipchat.com/#narrow/stream/116395-maths/topic/row.20rank.20equals.20column.20rank/near/350462992). -/ @@ -30,6 +35,8 @@ namespace matrix open finite_dimensional variables {l m n o R : Type*} [m_fin : fintype m] [fintype n] [fintype o] + +section comm_ring variables [comm_ring R] /-- The rank of a matrix is the rank of its image. -/ @@ -98,8 +105,7 @@ begin rw [rank, rank, mul_vec_lin_submatrix, linear_map.range_comp, linear_map.range_comp, (show linear_map.fun_left R R e.symm = linear_equiv.fun_congr_left R R e.symm, from rfl), linear_equiv.range, submodule.map_top], - -- TODO: generalize `finite_dimensional.finrank_map_le` and use it here - exact finrank_le_finrank_of_rank_le_rank (lift_rank_map_le _ _) (rank_lt_aleph_0 _ _), + exact submodule.finrank_map_le _ _, end lemma rank_reindex [fintype m] (e₁ e₂ : m ≃ n) (A : matrix m m R) : @@ -154,4 +160,110 @@ lemma rank_eq_finrank_span_cols (A : matrix m n R) : A.rank = finrank R (submodule.span R (set.range Aᵀ)) := by rw [rank, matrix.range_mul_vec_lin] +end comm_ring + +/-! ### Lemmas about transpose and conjugate transpose + +This section contains lemmas about the rank of `matrix.transpose` and `matrix.conj_transpose`. + +Unfortunately the proofs are essentially duplicated between the two; `ℚ` is a linearly-ordered ring +but can't be a star-ordered ring, while `ℂ` is star-ordered (with `open_locale complex_order`) but +not linearly ordered. For now we don't prove the transpose case for `ℂ`. + +TODO: the lemmas `matrix.rank_transpose` and `matrix.rank_conj_transpose` current follow a short +proof that is a simple consequence of `matrix.rank_transpose_mul_self` and +`matrix.rank_conj_transpose_mul_self`. This proof pulls in unecessary assumptions on `R`, and should +be replaced with a proof that uses Gaussian reduction or argues via linear combinations. +-/ + +section star_ordered_field +variables [fintype m] [field R] [partial_order R] [star_ordered_ring R] + +lemma ker_mul_vec_lin_conj_transpose_mul_self (A : matrix m n R) : + linear_map.ker (Aᴴ ⬝ A).mul_vec_lin = linear_map.ker (mul_vec_lin A):= +begin + ext x, + simp only [linear_map.mem_ker, mul_vec_lin_apply, ←mul_vec_mul_vec], + split, + { intro h, + replace h := congr_arg (dot_product (star x)) h, + rwa [dot_product_mul_vec, dot_product_zero, vec_mul_conj_transpose, star_star, + dot_product_star_self_eq_zero] at h }, + { intro h, rw [h, mul_vec_zero] }, +end + +lemma rank_conj_transpose_mul_self (A : matrix m n R) : + (Aᴴ ⬝ A).rank = A.rank := +begin + dunfold rank, + refine add_left_injective (finrank R (A.mul_vec_lin).ker) _, + dsimp only, + rw [linear_map.finrank_range_add_finrank_ker, + ←((Aᴴ ⬝ A).mul_vec_lin).finrank_range_add_finrank_ker], + congr' 1, + rw ker_mul_vec_lin_conj_transpose_mul_self, +end + +-- this follows the proof here https://math.stackexchange.com/a/81903/1896 +/-- TODO: prove this in greater generality. -/ +@[simp] lemma rank_conj_transpose (A : matrix m n R) : Aᴴ.rank = A.rank := +le_antisymm + (((rank_conj_transpose_mul_self _).symm.trans_le $ rank_mul_le_left _ _).trans_eq $ + congr_arg _ $ conj_transpose_conj_transpose _) + ((rank_conj_transpose_mul_self _).symm.trans_le $ rank_mul_le_left _ _) + +@[simp] lemma rank_self_mul_conj_transpose (A : matrix m n R) : (A ⬝ Aᴴ).rank = A.rank := +by simpa only [rank_conj_transpose, conj_transpose_conj_transpose] + using rank_conj_transpose_mul_self Aᴴ + +end star_ordered_field + +section linear_ordered_field +variables [fintype m] [linear_ordered_field R] + +lemma ker_mul_vec_lin_transpose_mul_self (A : matrix m n R) : + linear_map.ker (Aᵀ ⬝ A).mul_vec_lin = linear_map.ker (mul_vec_lin A):= +begin + ext x, + simp only [linear_map.mem_ker, mul_vec_lin_apply, ←mul_vec_mul_vec], + split, + { intro h, + replace h := congr_arg (dot_product x) h, + rwa [dot_product_mul_vec, dot_product_zero, vec_mul_transpose, + dot_product_self_eq_zero] at h }, + { intro h, rw [h, mul_vec_zero] }, +end + +lemma rank_transpose_mul_self (A : matrix m n R) : (Aᵀ ⬝ A).rank = A.rank := +begin + dunfold rank, + refine add_left_injective (finrank R (A.mul_vec_lin).ker) _, + dsimp only, + rw [linear_map.finrank_range_add_finrank_ker, + ←((Aᵀ ⬝ A).mul_vec_lin).finrank_range_add_finrank_ker], + congr' 1, + rw ker_mul_vec_lin_transpose_mul_self, +end + +/-- TODO: prove this in greater generality. -/ +@[simp] lemma rank_transpose (A : matrix m n R) : Aᵀ.rank = A.rank := +le_antisymm + ((rank_transpose_mul_self _).symm.trans_le $ rank_mul_le_left _ _) + ((rank_transpose_mul_self _).symm.trans_le $ rank_mul_le_left _ _) + +@[simp] lemma rank_self_mul_transpose (A : matrix m n R) : (A ⬝ Aᵀ).rank = A.rank := +by simpa only [rank_transpose, transpose_transpose] using rank_transpose_mul_self Aᵀ + +end linear_ordered_field + +/-- The rank of a matrix is the rank of the space spanned by its rows. + +TODO: prove this in a generality that works for `ℂ` too, not just `ℚ` and `ℝ`. -/ +lemma rank_eq_finrank_span_row [linear_ordered_field R] [finite m] (A : matrix m n R) : + A.rank = finrank R (submodule.span R (set.range A)) := +begin + casesI nonempty_fintype m, + rw [←rank_transpose, rank_eq_finrank_span_cols, transpose_transpose] +end + end matrix From 730c6d4cab72b9d84fcfb9e95e8796e9cd8f40ba Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Wed, 19 Apr 2023 18:47:19 +0000 Subject: [PATCH 0843/1291] chore(order/initial_seg): tweak `subsingleton_of_trichotomous_of_irrefl` (#18749) We rename it, turn it into an instance, and golf the next instance with it. --- src/order/initial_seg.lean | 11 +++++------ 1 file changed, 5 insertions(+), 6 deletions(-) diff --git a/src/order/initial_seg.lean b/src/order/initial_seg.lean index 6649efe760bb1..2d3ef20adcc9d 100644 --- a/src/order/initial_seg.lean +++ b/src/order/initial_seg.lean @@ -107,19 +107,18 @@ end⟩ @[simp] theorem trans_apply (f : r ≼i s) (g : s ≼i t) (a : α) : (f.trans g) a = g (f a) := rfl -theorem unique_of_trichotomous_of_irrefl [is_trichotomous β s] [is_irrefl β s] : - well_founded r → subsingleton (r ≼i s) | ⟨h⟩ := +instance subsingleton_of_trichotomous_of_irrefl [is_trichotomous β s] [is_irrefl β s] + [is_well_founded α r] : subsingleton (r ≼i s) := ⟨λ f g, begin ext a, - have := h a, induction this with a H IH, + apply is_well_founded.induction r a (λ b IH, _), refine extensional_of_trichotomous_of_irrefl s (λ x, _), - simp only [f.init_iff, g.init_iff], + rw [f.init_iff, g.init_iff], exact exists_congr (λ x, and_congr_left $ λ hx, IH _ hx ▸ iff.rfl) end⟩ instance [is_well_order β s] : subsingleton (r ≼i s) := -⟨λ a, @subsingleton.elim _ (unique_of_trichotomous_of_irrefl - (@rel_embedding.well_founded _ _ r s a is_well_founded.wf)) a⟩ +⟨λ a, by { letI := a.is_well_founded, apply subsingleton.elim }⟩ protected theorem eq [is_well_order β s] (f g : r ≼i s) (a) : f a = g a := by rw subsingleton.elim f g From 13b8e258f14bffb5def542aa78b803b0b80541aa Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 19 Apr 2023 21:06:45 +0000 Subject: [PATCH 0844/1291] feat(group_theory/monoid_localization): Order (#18724) Prove that every (linearly) ordered cancellative monoid can be embedded into a (linearly) ordered group, namely its Grothendieck group. Note that cancellativity is necessary since submonoids of a group are cancellative. Co-authored-by: Eric Wieser --- src/group_theory/monoid_localization.lean | 136 +++++++++++++++++++++- 1 file changed, 132 insertions(+), 4 deletions(-) diff --git a/src/group_theory/monoid_localization.lean b/src/group_theory/monoid_localization.lean index 20c82192d58bb..2c763e0eefe69 100644 --- a/src/group_theory/monoid_localization.lean +++ b/src/group_theory/monoid_localization.lean @@ -41,6 +41,9 @@ This defines the localization as a quotient type, `localization`, but the majori subsequent lemmas in the file are given in terms of localizations up to isomorphism, using maps which satisfy the characteristic predicate. +The Grothendieck group construction corresponds to localizing at the top submonoid, namely making +every element invertible. + ## Implementation notes In maths it is natural to reason up to isomorphism, but in Lean we cannot naturally `rewrite` one @@ -61,11 +64,20 @@ localization as a quotient type satisfies the characteristic predicate). The lem `mk_eq_monoid_of_mk'` hence gives you access to the results in the rest of the file, which are about the `localization_map.mk'` induced by any localization map. +## TODO + +* Show that the localization at the top monoid is a group. +* Generalise to (nonempty) subsemigroups. +* If we acquire more bundlings, we can make `localization.mk_order_embedding` be an ordered monoid + embedding. + ## Tags localization, monoid localization, quotient monoid, congruence relation, characteristic predicate, -commutative monoid +commutative monoid, grothendieck group -/ +open function + namespace add_submonoid variables {M : Type*} [add_comm_monoid M] (S : add_submonoid M) (N : Type*) [add_comm_monoid N] @@ -245,11 +257,11 @@ def mk (x : M) (y : S) : localization S := (r S).mk' (x, y) universes u /-- Dependent recursion principle for localizations: given elements `f a b : p (mk a b)` -for all `a b`, such that `r S (a, b) (c, d)` implies `f a b = f c d` (wih the correct coercions), +for all `a b`, such that `r S (a, b) (c, d)` implies `f a b = f c d` (with the correct coercions), then `f` is defined on the whole `localization S`. -/ @[elab_as_eliminator, to_additive -"Dependent recursion principle for `add_localizations`: given elements `f a b : p (mk a b)` -for all `a b`, such that `r S (a, b) (c, d)` implies `f a b = f c d` (wih the correct coercions), +"Dependent recursion principle for `add_localization`s: given elements `f a b : p (mk a b)` +for all `a b`, such that `r S (a, b) (c, d)` implies `f a b = f c d` (with the correct coercions), then `f` is defined on the whole `add_localization S`."] def rec {p : localization S → Sort u} (f : ∀ (a : M) (b : S), p (mk a b)) @@ -259,6 +271,16 @@ def rec {p : localization S → Sort u} quot.rec (λ y, eq.rec (f y.1 y.2) (prod.mk.eta : (y.1, y.2) = y)) (λ y z h, by { cases y, cases z, exact H h }) x +/-- Copy of `quotient.rec_on_subsingleton₂` for `localization` -/ +@[elab_as_eliminator, to_additive "Copy of `quotient.rec_on_subsingleton₂` for `add_localization`"] +def rec_on_subsingleton₂ {r : localization S → localization S → Sort u} + [h : ∀ (a c : M) (b d : S), subsingleton (r (mk a b) (mk c d))] + (x y : localization S) + (f : Π (a c : M) (b d : S), r (mk a b) (mk c d)) : r x y := +@quotient.rec_on_subsingleton₂' _ _ _ _ r + (prod.rec $ by exact λ _ _, prod.rec $ by exact λ _ _, h _ _ _ _) x y + (prod.rec $ by exact λ _ _, prod.rec $ by exact λ _ _, f _ _ _ _) + attribute [irreducible] localization @[to_additive] lemma mk_mul (a c : M) (b d : S) : mk a b * mk c d = mk (a * c) (b * d) := rfl @@ -1431,3 +1453,109 @@ end localization_with_zero_map end submonoid end comm_monoid_with_zero + +namespace localization +variables {α : Type*} [cancel_comm_monoid α] {s : submonoid α} {a₁ b₁ : α} {a₂ b₂ : s} + +@[to_additive] lemma mk_left_injective (b : s) : injective (λ a, mk a b) := +λ c d h, by simpa [-mk_eq_monoid_of_mk', mk_eq_mk_iff, r_iff_exists] using h + +@[to_additive] lemma mk_eq_mk_iff' : mk a₁ a₂ = mk b₁ b₂ ↔ ↑b₂ * a₁ = a₂ * b₁ := +by simp_rw [mk_eq_mk_iff, r_iff_exists, mul_left_cancel_iff, exists_const] + +@[to_additive] instance decidable_eq [decidable_eq α] : decidable_eq (localization s) := +λ a b, localization.rec_on_subsingleton₂ a b $ λ a₁ a₂ b₁ b₂, decidable_of_iff' _ mk_eq_mk_iff' + +end localization + +/-! ### Order -/ + +namespace localization +variables {α : Type*} + +section ordered_cancel_comm_monoid +variables [ordered_cancel_comm_monoid α] {s : submonoid α} {a₁ b₁ : α} {a₂ b₂ : s} + +@[to_additive] instance : has_le (localization s) := +⟨λ a b, localization.lift_on₂ a b (λ a₁ a₂ b₁ b₂, ↑b₂ * a₁ ≤ a₂ * b₁) $ + λ a₁ b₁ a₂ b₂ c₁ d₁ c₂ d₂ hab hcd, propext begin + obtain ⟨e, he⟩ := r_iff_exists.1 hab, + obtain ⟨f, hf⟩ := r_iff_exists.1 hcd, + simp only [mul_right_inj] at he hf, + dsimp, + rw [←mul_le_mul_iff_right, mul_right_comm, ←hf, mul_right_comm, mul_right_comm ↑a₂, + mul_le_mul_iff_right, ←mul_le_mul_iff_left, mul_left_comm, he, mul_left_comm, + mul_left_comm ↑b₂, mul_le_mul_iff_left], + end⟩ + +@[to_additive] instance : has_lt (localization s) := +⟨λ a b, localization.lift_on₂ a b (λ a₁ a₂ b₁ b₂, ↑b₂ * a₁ < a₂ * b₁) $ + λ a₁ b₁ a₂ b₂ c₁ d₁ c₂ d₂ hab hcd, propext begin + obtain ⟨e, he⟩ := r_iff_exists.1 hab, + obtain ⟨f, hf⟩ := r_iff_exists.1 hcd, + simp only [mul_right_inj] at he hf, + dsimp, + rw [←mul_lt_mul_iff_right, mul_right_comm, ←hf, mul_right_comm, mul_right_comm ↑a₂, + mul_lt_mul_iff_right, ←mul_lt_mul_iff_left, mul_left_comm, he, mul_left_comm, + mul_left_comm ↑b₂, mul_lt_mul_iff_left], + end⟩ + +@[to_additive] lemma mk_le_mk : mk a₁ a₂ ≤ mk b₁ b₂ ↔ ↑b₂ * a₁ ≤ a₂ * b₁ := iff.rfl +@[to_additive] lemma mk_lt_mk : mk a₁ a₂ < mk b₁ b₂ ↔ ↑b₂ * a₁ < a₂ * b₁ := iff.rfl + +@[to_additive] instance : ordered_cancel_comm_monoid (localization s) := +{ le_refl := λ a, localization.induction_on a $ λ a, le_rfl, + le_trans := λ a b c, localization.induction_on₃ a b c $ λ a b c hab hbc, begin + simp only [mk_le_mk] at ⊢ hab hbc, + refine le_of_mul_le_mul_left' _, + { exact b.2 }, + rw [mul_left_comm], + refine (mul_le_mul_left' hab _).trans _, + rwa [mul_left_comm, mul_left_comm ↑b.2, mul_le_mul_iff_left], + end, + le_antisymm := λ a b, begin + induction a with a₁ a₂, + induction b with b₁ b₂, + simp_rw [mk_le_mk, mk_eq_mk_iff, r_iff_exists], + exact λ hab hba, ⟨1, by rw [hab.antisymm hba]⟩, + all_goals { intros, refl }, + end, + lt_iff_le_not_le := λ a b, localization.induction_on₂ a b $ λ a b, lt_iff_le_not_le, + mul_le_mul_left := λ a b, localization.induction_on₂ a b $ λ a b hab c, + localization.induction_on c $ λ c, begin + simp only [mk_mul, mk_le_mk, submonoid.coe_mul, mul_mul_mul_comm _ _ c.1] at ⊢ hab, + exact mul_le_mul_left' hab _, + end, + le_of_mul_le_mul_left := λ a b c, localization.induction_on₃ a b c $ λ a b c hab, begin + simp only [mk_mul, mk_le_mk, submonoid.coe_mul, mul_mul_mul_comm _ _ a.1] at ⊢ hab, + exact le_of_mul_le_mul_left' hab, + end, + ..localization.comm_monoid _, ..localization.has_le, ..localization.has_lt } + +@[to_additive] instance decidable_le [decidable_rel ((≤) : α → α → Prop)] : + decidable_rel ((≤) : localization s → localization s → Prop) := +λ a b, localization.rec_on_subsingleton₂ a b $ λ a₁ a₂ b₁ b₂, decidable_of_iff' _ mk_le_mk + +@[to_additive] instance decidable_lt [decidable_rel ((<) : α → α → Prop)] : + decidable_rel ((<) : localization s → localization s → Prop) := +λ a b, localization.rec_on_subsingleton₂ a b $ λ a₁ a₂ b₁ b₂, decidable_of_iff' _ mk_lt_mk + +/-- An ordered cancellative monoid injects into its localization by sending `a` to `a / b`. -/ +@[to_additive "An ordered cancellative monoid injects into its localization by sending `a` to +`a - b`.", simps] def mk_order_embedding (b : s) : α ↪o localization s := +{ to_fun := λ a, mk a b, + inj' := mk_left_injective _, + map_rel_iff' := λ a b, by simp [-mk_eq_monoid_of_mk', mk_le_mk] } + +end ordered_cancel_comm_monoid + +@[to_additive] instance [linear_ordered_cancel_comm_monoid α] {s : submonoid α} : + linear_ordered_cancel_comm_monoid (localization s) := +{ le_total := λ a b, localization.induction_on₂ a b $ λ _ _, + by { simp_rw mk_le_mk, exact le_total _ _ }, + decidable_le := @localization.decidable_le α _ _ has_le.le.decidable, + decidable_lt := @localization.decidable_lt α _ _ has_lt.lt.decidable, + decidable_lt := localization.decidable_eq, + ..localization.ordered_cancel_comm_monoid } + +end localization From 9ac7c0c8c4d7a535ec3e5b34b8859aab9233b2f4 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 19 Apr 2023 23:18:38 +0000 Subject: [PATCH 0845/1291] feat(data/finset/basic): `insert` and `erase` lemmas (#18729) Interaction of `insert` and `erase` with `inter`, `union` and `disjoint`. Co-authored-by: Eric Rodriguez --- src/data/finset/basic.lean | 77 ++++++++++++++++++++++++++++++---- src/data/set/basic.lean | 5 +++ src/order/boolean_algebra.lean | 9 ++++ src/order/heyting/basic.lean | 8 ++++ 4 files changed, 91 insertions(+), 8 deletions(-) diff --git a/src/data/finset/basic.lean b/src/data/finset/basic.lean index 28c672b401f2e..30687beee91e0 100644 --- a/src/data/finset/basic.lean +++ b/src/data/finset/basic.lean @@ -863,7 +863,7 @@ end insert /-! ### Lattice structure -/ section lattice -variables [decidable_eq α] {s t u v : finset α} {a b : α} +variables [decidable_eq α] {s s₁ s₂ t t₁ t₂ u v : finset α} {a b : α} /-- `s ∪ t` is the set such that `a ∈ s ∪ t` iff `a ∈ s` or `a ∈ t`. -/ instance : has_union (finset α) := ⟨λ s t, ⟨_, t.2.ndunion s.1⟩⟩ @@ -918,6 +918,9 @@ theorem subset_union_right (s₁ s₂ : finset α) : s₂ ⊆ s₁ ∪ s₂ := lemma union_subset_union (hsu : s ⊆ u) (htv : t ⊆ v) : s ∪ t ⊆ u ∪ v := sup_le_sup (le_iff_subset.2 hsu) htv +lemma union_subset_union_left (h : s₁ ⊆ s₂) : s₁ ∪ t ⊆ s₂ ∪ t := union_subset_union h subset.rfl +lemma union_subset_union_right (h : t₁ ⊆ t₂) : s ∪ t₁ ⊆ s ∪ t₂ := union_subset_union subset.rfl h + lemma union_comm (s₁ s₂ : finset α) : s₁ ∪ s₂ = s₂ ∪ s₁ := sup_comm instance : is_commutative (finset α) (∪) := ⟨union_comm⟩ @@ -1396,6 +1399,9 @@ set.ext $ λ _, mem_sdiff lemma union_sdiff_left (s t : finset α) : (s ∪ t) \ s = t \ s := sup_sdiff_left_self lemma union_sdiff_right (s t : finset α) : (s ∪ t) \ t = s \ t := sup_sdiff_right_self +lemma union_sdiff_cancel_left (h : disjoint s t) : (s ∪ t) \ s = t := h.sup_sdiff_cancel_left +lemma union_sdiff_cancel_right (h : disjoint s t) : (s ∪ t) \ t = s := h.sup_sdiff_cancel_right + lemma union_sdiff_symm : s ∪ (t \ s) = t ∪ (s \ t) := by simp [union_comm] lemma sdiff_union_inter (s t : finset α) : (s \ t) ∪ (s ∩ t) = s := sup_sdiff_inf _ _ @@ -1444,15 +1450,69 @@ lemma sdiff_union_distrib (s t₁ t₂ : finset α) : s \ (t₁ ∪ t₂) = (s \ lemma union_sdiff_self (s t : finset α) : (s ∪ t) \ t = s \ t := sup_sdiff_right_self +-- TODO: Do we want to delete this lemma and `finset.disj_union_singleton`, +-- or instead add `finset.union_singleton`/`finset.singleton_union`? lemma sdiff_singleton_eq_erase (a : α) (s : finset α) : s \ singleton a = erase s a := by { ext, rw [mem_erase, mem_sdiff, mem_singleton], tauto } -@[simp] lemma sdiff_singleton_not_mem_eq_self (s : finset α) {a : α} (ha : a ∉ s) : s \ {a} = s := -by simp only [sdiff_singleton_eq_erase, ha, erase_eq_of_not_mem, not_false_iff] +-- This lemma matches `finset.insert_eq` in functionality. +lemma erase_eq (s : finset α) (a : α) : s.erase a = s \ {a} := (sdiff_singleton_eq_erase _ _).symm + +lemma disjoint_erase_comm : disjoint (s.erase a) t ↔ disjoint s (t.erase a) := +by simp_rw [erase_eq, disjoint_sdiff_comm] + +lemma disjoint_of_erase_left (ha : a ∉ t) (hst : disjoint (s.erase a) t) : disjoint s t := +by { rw [←erase_insert ha, ←disjoint_erase_comm, disjoint_insert_right], + exact ⟨not_mem_erase _ _, hst⟩ } + +lemma disjoint_of_erase_right (ha : a ∉ s) (hst : disjoint s (t.erase a)) : disjoint s t := +by { rw [←erase_insert ha, disjoint_erase_comm, disjoint_insert_left], + exact ⟨not_mem_erase _ _, hst⟩ } + +@[simp] lemma inter_erase (a : α) (s t : finset α) : s ∩ t.erase a = (s ∩ t).erase a := +by simp only [erase_eq, inter_sdiff] + +@[simp] lemma erase_inter (a : α) (s t : finset α) : s.erase a ∩ t = (s ∩ t).erase a := +by simpa only [inter_comm t] using inter_erase a t s + +lemma erase_sdiff_comm (s t : finset α) (a : α) : s.erase a \ t = (s \ t).erase a := +by simp_rw [erase_eq, sdiff_right_comm] + +lemma insert_union_comm (s t : finset α) (a : α) : insert a s ∪ t = s ∪ insert a t := +by rw [insert_union, union_insert] + +lemma erase_inter_comm (s t : finset α) (a : α) : s.erase a ∩ t = s ∩ t.erase a := +by rw [erase_inter, inter_erase] + +lemma erase_union_distrib (s t : finset α) (a : α) : (s ∪ t).erase a = s.erase a ∪ t.erase a := +by simp_rw [erase_eq, union_sdiff_distrib] + +lemma insert_inter_distrib (s t : finset α) (a : α) : insert a (s ∩ t) = insert a s ∩ insert a t := +by simp_rw [insert_eq, union_distrib_left] + +lemma erase_sdiff_distrib (s t : finset α) (a : α) : (s \ t).erase a = s.erase a \ t.erase a := +by simp_rw [erase_eq, sdiff_sdiff, sup_sdiff_eq_sup le_rfl, sup_comm] + +lemma erase_union_of_mem (ha : a ∈ t) (s : finset α) : s.erase a ∪ t = s ∪ t := +by rw [←insert_erase (mem_union_right s ha), erase_union_distrib, ←union_insert, insert_erase ha] + +lemma union_erase_of_mem (ha : a ∈ s) (t : finset α) : s ∪ t.erase a = s ∪ t := +by rw [←insert_erase (mem_union_left t ha), erase_union_distrib, ←insert_union, insert_erase ha] + +@[simp] lemma sdiff_singleton_eq_self (ha : a ∉ s) : s \ {a} = s := +sdiff_eq_self_iff_disjoint.2 $ by simp [ha] lemma sdiff_sdiff_left' (s t u : finset α) : (s \ t) \ u = (s \ t) ∩ (s \ u) := sdiff_sdiff_left' +lemma sdiff_union_sdiff_cancel (hts : t ⊆ s) (hut : u ⊆ t) : s \ t ∪ t \ u = s \ u := +sdiff_sup_sdiff_cancel hts hut + +lemma sdiff_union_erase_cancel (hts : t ⊆ s) (ha : a ∈ t) : s \ t ∪ t.erase a = s.erase a := +by simp_rw [erase_eq, sdiff_union_sdiff_cancel hts (singleton_subset_iff.2 ha)] + +lemma sdiff_sdiff_eq_sdiff_union (h : u ⊆ s) : s \ (t \ u) = s \ t ∪ u := sdiff_sdiff_eq_sdiff_sup h + lemma sdiff_insert (s t : finset α) (x : α) : s \ insert x t = (s \ t).erase x := by simp_rw [← sdiff_singleton_eq_erase, insert_eq, @@ -1462,11 +1522,12 @@ lemma sdiff_insert_insert_of_mem_of_not_mem {s t : finset α} {x : α} (hxs : x insert x (s \ insert x t) = s \ t := by rw [sdiff_insert, insert_erase (mem_sdiff.mpr ⟨hxs, hxt⟩)] -lemma sdiff_erase {x : α} (hx : x ∈ s) : s \ s.erase x = {x} := -begin - rw [← sdiff_singleton_eq_erase, sdiff_sdiff_right_self], - exact inf_eq_right.2 (singleton_subset_iff.2 hx), -end +lemma sdiff_erase (h : a ∈ s) : s \ t.erase a = insert a (s \ t) := +by rw [←sdiff_singleton_eq_erase, sdiff_sdiff_eq_sdiff_union (singleton_subset_iff.2 h), insert_eq, + union_comm] + +lemma sdiff_erase_self (ha : a ∈ s) : s \ s.erase a = {a} := +by rw [sdiff_erase ha, sdiff_self, insert_emptyc_eq] lemma sdiff_sdiff_self_left (s t : finset α) : s \ (s \ t) = s ∩ t := sdiff_sdiff_right_self diff --git a/src/data/set/basic.lean b/src/data/set/basic.lean index e6224ffb184f7..08af3259a3f2e 100644 --- a/src/data/set/basic.lean +++ b/src/data/set/basic.lean @@ -960,6 +960,11 @@ disjoint_sup_right lemma disjoint_sdiff_left : disjoint (t \ s) s := disjoint_sdiff_self_left lemma disjoint_sdiff_right : disjoint s (t \ s) := disjoint_sdiff_self_right +lemma diff_union_diff_cancel (hts : t ⊆ s) (hut : u ⊆ t) : s \ t ∪ t \ u = s \ u := +sdiff_sup_sdiff_cancel hts hut + +lemma diff_diff_eq_sdiff_union (h : u ⊆ s) : s \ (t \ u) = s \ t ∪ u := sdiff_sdiff_eq_sdiff_sup h + @[simp] lemma disjoint_singleton_left : disjoint {a} s ↔ a ∉ s := by simp [set.disjoint_iff, subset_def]; exact iff.rfl diff --git a/src/order/boolean_algebra.lean b/src/order/boolean_algebra.lean index b83a09204696b..b6772b7ac56b7 100644 --- a/src/order/boolean_algebra.lean +++ b/src/order/boolean_algebra.lean @@ -385,6 +385,9 @@ by rw [sdiff_inf, sdiff_eq_bot_iff.2 inf_le_left, bot_sup_eq, inf_sdiff_assoc] lemma inf_sdiff_distrib_right (a b c : α) : a \ b ⊓ c = (a ⊓ c) \ (b ⊓ c) := by simp_rw [@inf_comm _ _ _ c, inf_sdiff_distrib_left] +lemma disjoint_sdiff_comm : disjoint (x \ z) y ↔ disjoint x (y \ z) := +by simp_rw [disjoint_iff, inf_sdiff_right_comm, inf_sdiff_assoc] + lemma sup_eq_sdiff_sup_sdiff_sup_inf : x ⊔ y = (x \ y) ⊔ (y \ x) ⊔ (x ⊓ y) := eq.symm $ calc (x \ y) ⊔ (y \ x) ⊔ (x ⊓ y) = @@ -569,6 +572,12 @@ by rw [←le_compl_iff_disjoint_left, compl_compl] lemma disjoint_compl_right_iff : disjoint x yᶜ ↔ x ≤ y := by rw [←le_compl_iff_disjoint_right, compl_compl] +lemma codisjoint_himp_self_left : codisjoint (x ⇨ y) x := @disjoint_sdiff_self_left αᵒᵈ _ _ _ +lemma codisjoint_himp_self_right : codisjoint x (x ⇨ y) := @disjoint_sdiff_self_right αᵒᵈ _ _ _ + +lemma himp_le : x ⇨ y ≤ z ↔ y ≤ z ∧ codisjoint x z := +(@le_sdiff αᵒᵈ _ _ _ _).trans $ and_congr_right' codisjoint.comm + end boolean_algebra instance Prop.boolean_algebra : boolean_algebra Prop := diff --git a/src/order/heyting/basic.lean b/src/order/heyting/basic.lean index 8d190745d9817..70dd95d2b7175 100644 --- a/src/order/heyting/basic.lean +++ b/src/order/heyting/basic.lean @@ -295,6 +295,10 @@ by rw [himp_inf_distrib, himp_self, top_inf_eq, h.himp_eq_left] lemma codisjoint.himp_inf_cancel_left (h : codisjoint a b) : b ⇨ (a ⊓ b) = a := by rw [himp_inf_distrib, himp_self, inf_top_eq, h.himp_eq_right] +/-- See `himp_le` for a stronger version in Boolean algebras. -/ +lemma codisjoint.himp_le_of_right_le (hac : codisjoint a c) (hba : b ≤ a) : c ⇨ b ≤ a := +(himp_le_himp_left hba).trans_eq hac.himp_eq_right + lemma le_himp_himp : a ≤ (a ⇨ b) ⇨ b := le_himp_iff.2 inf_himp_le lemma himp_triangle (a b c : α) : (a ⇨ b) ⊓ (b ⇨ c) ≤ a ⇨ c := @@ -442,6 +446,10 @@ by rw [sup_sdiff, sdiff_self, bot_sup_eq, h.sdiff_eq_right] lemma disjoint.sup_sdiff_cancel_right (h : disjoint a b) : (a ⊔ b) \ b = a := by rw [sup_sdiff, sdiff_self, sup_bot_eq, h.sdiff_eq_left] +/-- See `le_sdiff` for a stronger version in generalised Boolean algebras. -/ +lemma disjoint.le_sdiff_of_le_left (hac : disjoint a c) (hab : a ≤ b) : a ≤ b \ c := +hac.sdiff_eq_left.ge.trans $ sdiff_le_sdiff_right hab + lemma sdiff_sdiff_le : a \ (a \ b) ≤ b := sdiff_le_iff.2 le_sdiff_sup lemma sdiff_triangle (a b c : α) : a \ c ≤ a \ b ⊔ b \ c := From cdc34484a07418af43daf8198beaf5c00324bca8 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 20 Apr 2023 07:06:15 +0000 Subject: [PATCH 0846/1291] refactor(algebra/module/dedekind_domain): eliminate existentials (#14734) This extracts some constructions from some long existentials so that we don't have to prove everything about the construction in a single place. This makes some of the statements longer, so in places the existential version is kept in terms of the definition one. --- src/algebra/module/dedekind_domain.lean | 26 +++++++++++++------ src/algebra/module/pid.lean | 33 ++++++++++++++++++------- src/algebra/module/torsion.lean | 30 +++++++++++++--------- 3 files changed, 60 insertions(+), 29 deletions(-) diff --git a/src/algebra/module/dedekind_domain.lean b/src/algebra/module/dedekind_domain.lean index db6de92fae59f..eee85d418a999 100644 --- a/src/algebra/module/dedekind_domain.lean +++ b/src/algebra/module/dedekind_domain.lean @@ -26,17 +26,17 @@ namespace submodule variables [is_dedekind_domain R] open unique_factorization_monoid +open_locale classical + /--Over a Dedekind domain, a `I`-torsion module is the internal direct sum of its `p i ^ e i`- torsion submodules, where `I = ∏ i, p i ^ e i` is its unique decomposition in prime ideals.-/ lemma is_internal_prime_power_torsion_of_is_torsion_by_ideal {I : ideal R} (hI : I ≠ ⊥) (hM : module.is_torsion_by_set R M I) : - ∃ (P : finset $ ideal R) [decidable_eq P] [∀ p ∈ P, prime p] (e : P → ℕ), - by exactI direct_sum.is_internal (λ p : P, torsion_by_set R M (p ^ e p : ideal R)) := + direct_sum.is_internal (λ p : (factors I).to_finset, + torsion_by_set R M (p ^ (factors I).count p : ideal R)) := begin - classical, let P := factors I, have prime_of_mem := λ p (hp : p ∈ P.to_finset), prime_of_factor p (multiset.mem_to_finset.mp hp), - refine ⟨P.to_finset, infer_instance, prime_of_mem, λ i, P.count i, _⟩, apply @torsion_by_set_is_internal _ _ _ _ _ _ _ _ (λ p, p ^ P.count p) _, { convert hM, rw [← finset.inf_eq_infi, is_dedekind_domain.inf_prime_pow_eq_prod, @@ -55,15 +55,25 @@ begin end /--A finitely generated torsion module over a Dedekind domain is an internal direct sum of its -`p i ^ e i`-torsion submodules for some prime ideals `p i` and numbers `e i`.-/ +`p i ^ e i`-torsion submodules where `p i` are factors of `(⊤ : submodule R M).annihilator` and +`e i` are their multiplicities. -/ theorem is_internal_prime_power_torsion [module.finite R M] (hM : module.is_torsion R M) : - ∃ (P : finset $ ideal R) [decidable_eq P] [∀ p ∈ P, prime p] (e : P → ℕ), - by exactI direct_sum.is_internal (λ p : P, torsion_by_set R M (p ^ e p : ideal R)) := + direct_sum.is_internal (λ p : (factors (⊤ : submodule R M).annihilator).to_finset, + torsion_by_set R M (p ^ (factors (⊤ : submodule R M).annihilator).count p : ideal R)) := begin - obtain ⟨I, hI, hM'⟩ := is_torsion_by_ideal_of_finite_of_is_torsion hM, + have hM' := module.is_torsion_by_set_annihilator_top R M, + have hI := submodule.annihilator_top_inter_non_zero_divisors hM, refine is_internal_prime_power_torsion_of_is_torsion_by_ideal _ hM', rw ←set.nonempty_iff_ne_empty at hI, rw submodule.ne_bot_iff, obtain ⟨x, H, hx⟩ := hI, exact ⟨x, H, non_zero_divisors.ne_zero hx⟩ end +/--A finitely generated torsion module over a Dedekind domain is an internal direct sum of its +`p i ^ e i`-torsion submodules for some prime ideals `p i` and numbers `e i`.-/ +theorem exists_is_internal_prime_power_torsion [module.finite R M] (hM : module.is_torsion R M) : + ∃ (P : finset $ ideal R) [decidable_eq P] [∀ p ∈ P, prime p] (e : P → ℕ), + by exactI direct_sum.is_internal (λ p : P, torsion_by_set R M (p ^ e p : ideal R)) := +⟨_, _, λ p hp, prime_of_factor p (multiset.mem_to_finset.mp hp), _, + is_internal_prime_power_torsion hM⟩ + end submodule diff --git a/src/algebra/module/pid.lean b/src/algebra/module/pid.lean index f80a4305ff9c6..fe80596510b4a 100644 --- a/src/algebra/module/pid.lean +++ b/src/algebra/module/pid.lean @@ -47,7 +47,7 @@ Finitely generated module, principal ideal domain, classification, structure the -/ universes u v -open_locale big_operators +open_locale big_operators classical variables {R : Type u} [comm_ring R] [is_domain R] [is_principal_ideal_ring R] variables {M : Type v} [add_comm_group M] [module R M] @@ -56,21 +56,36 @@ variables {N : Type (max u v)} [add_comm_group N] [module R N] open_locale direct_sum open submodule +open unique_factorization_monoid + /--A finitely generated torsion module over a PID is an internal direct sum of its `p i ^ e i`-torsion submodules for some primes `p i` and numbers `e i`.-/ theorem submodule.is_internal_prime_power_torsion_of_pid + [module.finite R M] (hM : module.is_torsion R M) : + direct_sum.is_internal (λ p : (factors (⊤ : submodule R M).annihilator).to_finset, + torsion_by R M + (is_principal.generator (p : ideal R) + ^ (factors (⊤ : submodule R M).annihilator).count p)) := +begin + convert is_internal_prime_power_torsion hM, + ext p : 1, + rw [← torsion_by_span_singleton_eq, ideal.submodule_span_eq, ← ideal.span_singleton_pow, + ideal.span_singleton_generator], +end + +/--A finitely generated torsion module over a PID is an internal direct sum of its +`p i ^ e i`-torsion submodules for some primes `p i` and numbers `e i`.-/ +theorem submodule.exists_is_internal_prime_power_torsion_of_pid [module.finite R M] (hM : module.is_torsion R M) : ∃ (ι : Type u) [fintype ι] [decidable_eq ι] (p : ι → R) (h : ∀ i, irreducible $ p i) (e : ι → ℕ), by exactI direct_sum.is_internal (λ i, torsion_by R M $ p i ^ e i) := begin - obtain ⟨P, dec, hP, e, this⟩ := is_internal_prime_power_torsion hM, - refine ⟨P, infer_instance, dec, λ p, is_principal.generator (p : ideal R), _, e, _⟩, + refine ⟨_, _, _, _, _, _, submodule.is_internal_prime_power_torsion_of_pid hM⟩, + exact finset.fintype_coe_sort _, { rintro ⟨p, hp⟩, - haveI := ideal.is_prime_of_prime (hP p hp), - exact (is_principal.prime_generator_of_is_prime p (hP p hp).ne_zero).irreducible }, - { convert this, ext p : 1, - rw [← torsion_by_span_singleton_eq, ideal.submodule_span_eq, ← ideal.span_singleton_pow, - ideal.span_singleton_generator] } + have hP := prime_of_factor p (multiset.mem_to_finset.mp hp), + haveI := ideal.is_prime_of_prime hP, + exact (is_principal.prime_generator_of_is_prime p hP.ne_zero).irreducible }, end namespace module @@ -205,7 +220,7 @@ theorem equiv_direct_sum_of_is_torsion [h' : module.finite R N] (hN : module.is_ ∃ (ι : Type u) [fintype ι] (p : ι → R) (h : ∀ i, irreducible $ p i) (e : ι → ℕ), nonempty $ N ≃ₗ[R] ⨁ (i : ι), R ⧸ R ∙ (p i ^ e i) := begin - obtain ⟨I, fI, _, p, hp, e, h⟩ := submodule.is_internal_prime_power_torsion_of_pid hN, + obtain ⟨I, fI, _, p, hp, e, h⟩ := submodule.exists_is_internal_prime_power_torsion_of_pid hN, haveI := fI, have : ∀ i, ∃ (d : ℕ) (k : fin d → ℕ), nonempty $ torsion_by R N (p i ^ e i) ≃ₗ[R] ⨁ j, R ⧸ R ∙ (p i ^ k j), diff --git a/src/algebra/module/torsion.lean b/src/algebra/module/torsion.lean index 2921d1496bda2..5b2743ad7dcaa 100644 --- a/src/algebra/module/torsion.lean +++ b/src/algebra/module/torsion.lean @@ -496,19 +496,25 @@ section torsion variables [comm_semiring R] [add_comm_monoid M] [module R M] open_locale big_operators -lemma is_torsion_by_ideal_of_finite_of_is_torsion [module.finite R M] (hM : module.is_torsion R M) : - ∃ I : ideal R, (I : set R) ∩ R⁰ ≠ ∅ ∧ module.is_torsion_by_set R M I := +variables (R M) + +lemma _root_.module.is_torsion_by_set_annihilator_top : + module.is_torsion_by_set R M (⊤ : submodule R M).annihilator := +λ x ha, mem_annihilator.mp ha.prop x mem_top + +variables {R M} + +lemma _root_.submodule.annihilator_top_inter_non_zero_divisors [module.finite R M] + (hM : module.is_torsion R M) : + ((⊤ : submodule R M).annihilator : set R) ∩ R⁰ ≠ ∅:= begin - cases (module.finite_def.mp infer_instance : (⊤ : submodule R M).fg) with S h, - refine ⟨∏ x in S, ideal.torsion_of R M x, _, _⟩, - { refine set.nonempty.ne_empty ⟨_, _, (∏ x in S, (@hM x).some : R⁰).2⟩, - rw [subtype.val_eq_coe, submonoid.coe_finset_prod], - apply ideal.prod_mem_prod, - exact λ x _, (@hM x).some_spec }, - { rw [module.is_torsion_by_set_iff_torsion_by_set_eq_top, eq_top_iff, ← h, span_le], - intros x hx, apply torsion_by_set_le_torsion_by_set_of_subset, - { apply ideal.le_of_dvd, exact finset.dvd_prod_of_mem _ hx }, - { rw mem_torsion_by_set_iff, rintro ⟨a, ha⟩, exact ha } } + obtain ⟨S, hS⟩ := ‹module.finite R M›.out, + refine set.nonempty.ne_empty ⟨_, _, (∏ x in S, (@hM x).some : R⁰).prop⟩, + rw [submonoid.coe_finset_prod, set_like.mem_coe, ←hS, mem_annihilator_span], + intro n, + letI := classical.dec_eq M, + rw [←finset.prod_erase_mul _ _ n.prop, mul_smul, ←submonoid.smul_def, (@hM n).some_spec, + smul_zero], end variables [no_zero_divisors R] [nontrivial R] From 4020ddee5b4580a409bfda7d2f42726ce86ae674 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 20 Apr 2023 08:38:43 +0000 Subject: [PATCH 0847/1291] feat(data/set/intervals/unordered_interval): `prod` and `pi` lemmas (#18835) --- src/data/set/intervals/pi.lean | 9 +++++++++ src/data/set/intervals/unordered_interval.lean | 12 +++++++++++- 2 files changed, 20 insertions(+), 1 deletion(-) diff --git a/src/data/set/intervals/pi.lean b/src/data/set/intervals/pi.lean index 1ce85fa64b8e2..3197184c38759 100644 --- a/src/data/set/intervals/pi.lean +++ b/src/data/set/intervals/pi.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import data.set.intervals.basic +import data.set.intervals.unordered_interval import data.set.lattice /-! @@ -105,6 +106,14 @@ end end pi_preorder +section pi_lattice +variables [Π i, lattice (α i)] + +@[simp] lemma pi_univ_uIcc (a b : Π i, α i) : pi univ (λ i, uIcc (a i) (b i)) = uIcc a b := +pi_univ_Icc _ _ + +end pi_lattice + variables [decidable_eq ι] [Π i, linear_order (α i)] open function (update) diff --git a/src/data/set/intervals/unordered_interval.lean b/src/data/set/intervals/unordered_interval.lean index 2bd90389e2153..1e795686afe08 100644 --- a/src/data/set/intervals/unordered_interval.lean +++ b/src/data/set/intervals/unordered_interval.lean @@ -41,7 +41,7 @@ variables {α β : Type*} namespace set section lattice -variables [lattice α] {a a₁ a₂ b b₁ b₂ c x : α} +variables [lattice α] [lattice β] {a a₁ a₂ b b₁ b₂ c x : α} /-- `uIcc a b` is the set of elements lying between `a` and `b`, with `a` and `b` included. Note that we define it more generally in a lattice as `set.Icc (a ⊓ b) (a ⊔ b)`. In a product type, @@ -96,6 +96,16 @@ lemma bdd_below_bdd_above_iff_subset_uIcc (s : set α) : bdd_below_bdd_above_iff_subset_Icc.trans ⟨λ ⟨a, b, h⟩, ⟨a, b, λ x hx, Icc_subset_uIcc (h hx)⟩, λ ⟨a, b, h⟩, ⟨_, _, h⟩⟩ +section prod + +@[simp] lemma uIcc_prod_uIcc (a₁ a₂ : α) (b₁ b₂ : β) : + [a₁, a₂] ×ˢ [b₁, b₂] = [(a₁, b₁), (a₂, b₂)] := +Icc_prod_Icc _ _ _ _ + +lemma uIcc_prod_eq (a b : α × β) : [a, b] = [a.1, b.1] ×ˢ [a.2, b.2] := by simp + +end prod + end lattice open_locale interval From 485b24ed47b1b7978d38a1e445158c6224c3f42c Mon Sep 17 00:00:00 2001 From: Christopher Hoskin Date: Thu, 20 Apr 2023 11:35:08 +0000 Subject: [PATCH 0848/1291] feat(order/directed): Scott continuous functions (#18517) We prove an insert result for directed sets when the relation is reflexive. This is then used to show that a Scott continuous function is monotone. This result is required in the [construction of the Scott topology on a preorder](https://github.com/leanprover-community/mathlib4/pull/2508) (see also #18448). Holding PR for mathlib4: leanprover-community/mathlib4#2543 Co-authored-by: Christopher Hoskin Co-authored-by: Eric Wieser --- src/order/directed.lean | 76 +++++++++++++++++++++++++++++++++++++++++ 1 file changed, 76 insertions(+) diff --git a/src/order/directed.lean b/src/order/directed.lean index 2bb601c877ab8..15957a24b5695 100644 --- a/src/order/directed.lean +++ b/src/order/directed.lean @@ -6,6 +6,7 @@ Authors: Johannes Hölzl import data.set.image import order.lattice import order.max +import order.bounds.basic /-! # Directed indexed families and sets @@ -22,6 +23,11 @@ directed iff each pair of elements has a shared upper bound. * `directed_on r s`: Predicate stating that the set `s` is `r`-directed. * `is_directed α r`: Prop-valued mixin stating that `α` is `r`-directed. Follows the style of the unbundled relation classes such as `is_total`. +* `scott_continuous`: Predicate stating that a function between preorders preserves + `is_lub` on directed sets. + +## References +* [Gierz et al, *A Compendium of Continuous Lattices*][GierzEtAl1980] -/ open function @@ -161,6 +167,37 @@ by assumption instance order_dual.is_directed_le [has_le α] [is_directed α (≥)] : is_directed αᵒᵈ (≤) := by assumption +section reflexive + +lemma directed_on.insert (h : reflexive r) (a : α) {s : set α} (hd : directed_on r s) + (ha : ∀ b ∈ s, ∃ c ∈ s, a ≼ c ∧ b ≼ c) : directed_on r (insert a s) := +begin + rintros x (rfl | hx) y (rfl | hy), + { exact ⟨y, set.mem_insert _ _, h _, h _⟩ }, + { obtain ⟨w, hws, hwr⟩ := ha y hy, + exact ⟨w, set.mem_insert_of_mem _ hws, hwr⟩ }, + { obtain ⟨w, hws, hwr⟩ := ha x hx, + exact ⟨w, set.mem_insert_of_mem _ hws, hwr.symm⟩ }, + { obtain ⟨w, hws, hwr⟩ := hd x hx y hy, + exact ⟨w, set.mem_insert_of_mem _ hws, hwr⟩ }, +end + +lemma directed_on_singleton (h : reflexive r) (a : α) : directed_on r ({a} : set α) := +λ x hx y hy, ⟨x, hx, h _, hx.symm ▸ hy.symm ▸ h _⟩ + +lemma directed_on_pair (h : reflexive r) {a b : α} (hab : a ≼ b) : + directed_on r ({a, b} : set α) := +(directed_on_singleton h _).insert h _ $ λ c hc, ⟨c, hc, hc.symm ▸ hab, h _⟩ + +lemma directed_on_pair' (h : reflexive r) {a b : α} (hab : a ≼ b) : + directed_on r ({b, a} : set α) := +begin + rw set.pair_comm, + apply directed_on_pair h hab, +end + +end reflexive + section preorder variables [preorder α] {a : α} @@ -218,3 +255,42 @@ instance order_top.to_is_directed_le [has_le α] [order_top α] : is_directed α @[priority 100] -- see Note [lower instance priority] instance order_bot.to_is_directed_ge [has_le α] [order_bot α] : is_directed α (≥) := ⟨λ a b, ⟨⊥, bot_le, bot_le⟩⟩ + +section scott_continuous + +variables [preorder α] {a : α} + +/-- +A function between preorders is said to be Scott continuous if it preserves `is_lub` on directed +sets. It can be shown that a function is Scott continuous if and only if it is continuous wrt the +Scott topology. + +The dual notion + +```lean +∀ ⦃d : set α⦄, d.nonempty → directed_on (≥) d → ∀ ⦃a⦄, is_glb d a → is_glb (f '' d) (f a) +``` + +does not appear to play a significant role in the literature, so is omitted here. +-/ +def scott_continuous [preorder β] (f : α → β) : Prop := +∀ ⦃d : set α⦄, d.nonempty → directed_on (≤) d → ∀ ⦃a⦄, is_lub d a → is_lub (f '' d) (f a) + +protected lemma scott_continuous.monotone [preorder β] {f : α → β} + (h : scott_continuous f) : + monotone f := +begin + intros a b hab, + have e1 : is_lub (f '' {a, b}) (f b), + { apply h, + { exact set.insert_nonempty _ _ }, + { exact directed_on_pair le_refl hab }, + { rw [is_lub, upper_bounds_insert, upper_bounds_singleton, + set.inter_eq_self_of_subset_right (set.Ici_subset_Ici.mpr hab)], + exact is_least_Ici } }, + apply e1.1, + rw set.image_pair, + exact set.mem_insert _ _, +end + +end scott_continuous From f784cc6142443d9ee623a20788c282112c322081 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 20 Apr 2023 11:35:10 +0000 Subject: [PATCH 0849/1291] feat(linear_algebra/tensor_product/matrix): connect with `matrix.kronecker` (#18616) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This shows that `to_matrix _ _ (tensor_product.map f g) = to_matrix _ _ f ⊗ₖ to_matrix _ _ g` , and the equivalent statement for `to_lin`. --- src/linear_algebra/tensor_product/matrix.lean | 78 +++++++++++++++++++ src/linear_algebra/tensor_product_basis.lean | 6 ++ 2 files changed, 84 insertions(+) create mode 100644 src/linear_algebra/tensor_product/matrix.lean diff --git a/src/linear_algebra/tensor_product/matrix.lean b/src/linear_algebra/tensor_product/matrix.lean new file mode 100644 index 0000000000000..a4490a52c0edc --- /dev/null +++ b/src/linear_algebra/tensor_product/matrix.lean @@ -0,0 +1,78 @@ +/- +Copyright (c) 2023 Eric Wieser. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Eric Wieser +-/ + +import data.matrix.kronecker +import linear_algebra.matrix.to_lin +import linear_algebra.tensor_product_basis + +/-! +# Connections between `tensor_product` and `matrix` + +This file contains results about the matrices corresponding to maps between tensor product types, +where the correspondance is induced by `basis.tensor_product` + +Notably, `tensor_product.to_matrix_map` shows that taking the tensor product of linear maps is +equivalent to taking the kronecker product of their matrix representations. +-/ + +variables {R : Type*} {M N P M' N' : Type*} {ι κ τ ι' κ' : Type*} +variables [decidable_eq ι] [decidable_eq κ] [decidable_eq τ] +variables [fintype ι] [fintype κ] [fintype τ] [fintype ι'] [fintype κ'] +variables [comm_ring R] +variables [add_comm_group M] [add_comm_group N] [add_comm_group P] +variables [add_comm_group M'] [add_comm_group N'] +variables [module R M] [module R N] [module R P] [module R M'] [module R N'] +variables (bM : basis ι R M) (bN : basis κ R N) (bP : basis τ R P) +variables (bM' : basis ι' R M') (bN' : basis κ' R N') + +open_locale kronecker +open matrix linear_map + +/-- The linear map built from `tensor_product.map` corresponds to the matrix built from +`matrix.kronecker`. -/ +lemma tensor_product.to_matrix_map (f : M →ₗ[R] M') (g : N →ₗ[R] N') : + to_matrix (bM.tensor_product bN) (bM'.tensor_product bN') (tensor_product.map f g) + = to_matrix bM bM' f ⊗ₖ to_matrix bN bN' g := +begin + ext ⟨i, j⟩ ⟨i', j'⟩, + simp_rw [matrix.kronecker_map_apply, to_matrix_apply, basis.tensor_product_apply, + tensor_product.map_tmul, basis.tensor_product_repr_tmul_apply], +end + +/-- The matrix built from `matrix.kronecker` corresponds to the linear map built from +`tensor_product.map`. -/ +lemma matrix.to_lin_kronecker (A : matrix ι' ι R) (B : matrix κ' κ R) : + to_lin (bM.tensor_product bN) (bM'.tensor_product bN') (A ⊗ₖ B) = + tensor_product.map (to_lin bM bM' A) (to_lin bN bN' B) := +by rw [←linear_equiv.eq_symm_apply, to_lin_symm, tensor_product.to_matrix_map, + to_matrix_to_lin, to_matrix_to_lin] + +/-- `tensor_product.comm` corresponds to a permutation of the identity matrix. -/ +lemma tensor_product.to_matrix_comm : + to_matrix (bM.tensor_product bN) (bN.tensor_product bM) (tensor_product.comm R M N) + = (1 : matrix (ι × κ) (ι × κ) R).submatrix prod.swap id := +begin + ext ⟨i, j⟩ ⟨i', j'⟩, + simp_rw [to_matrix_apply, basis.tensor_product_apply, linear_equiv.coe_coe, + tensor_product.comm_tmul, basis.tensor_product_repr_tmul_apply, matrix.submatrix_apply, + prod.swap_prod_mk, id.def, basis.repr_self_apply, matrix.one_apply, prod.ext_iff, ite_and, + @eq_comm _ i', @eq_comm _ j'], + split_ifs; simp, +end + +/-- `tensor_product.assoc` corresponds to a permutation of the identity matrix. -/ +lemma tensor_product.to_matrix_assoc : + to_matrix ((bM.tensor_product bN).tensor_product bP) (bM.tensor_product (bN.tensor_product bP)) + (tensor_product.assoc R M N P) + = (1 : matrix (ι × κ × τ) (ι × κ × τ) R).submatrix id (equiv.prod_assoc _ _ _) := +begin + ext ⟨i, j, k⟩ ⟨⟨i', j'⟩, k'⟩, + simp_rw [to_matrix_apply, basis.tensor_product_apply, linear_equiv.coe_coe, + tensor_product.assoc_tmul, basis.tensor_product_repr_tmul_apply, matrix.submatrix_apply, + equiv.prod_assoc_apply, id.def, basis.repr_self_apply, matrix.one_apply, prod.ext_iff, ite_and, + @eq_comm _ i', @eq_comm _ j', @eq_comm _ k'], + split_ifs; simp, +end diff --git a/src/linear_algebra/tensor_product_basis.lean b/src/linear_algebra/tensor_product_basis.lean index cc422c6c209ab..1fdf3948c96f4 100644 --- a/src/linear_algebra/tensor_product_basis.lean +++ b/src/linear_algebra/tensor_product_basis.lean @@ -40,4 +40,10 @@ lemma basis.tensor_product_apply' (b : basis ι R M) (c : basis κ R N) (i : ι basis.tensor_product b c i = b i.1 ⊗ₜ c i.2 := by simp [basis.tensor_product] +@[simp] +lemma basis.tensor_product_repr_tmul_apply (b : basis ι R M) (c : basis κ R N) + (m : M) (n : N) (i : ι) (j : κ) : + (basis.tensor_product b c).repr (m ⊗ₜ n) (i, j) = b.repr m i * c.repr n j := +by simp [basis.tensor_product] + end comm_ring From 44e29dbcff83ba7114a464d592b8c3743987c1e5 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 20 Apr 2023 11:35:11 +0000 Subject: [PATCH 0850/1291] feat(algebra/order/ring/defs): negative versions of order lemmas (#18831) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit For each of a handful of lemmas with a `0 <` or `0 ≤` assumption, this adds a variant with a `< 0` or `≤ 0` assumption. --- src/algebra/order/field/basic.lean | 3 +++ src/algebra/order/ring/defs.lean | 32 ++++++++++++++++++++++++++++++ src/algebra/order/ring/lemmas.lean | 5 ++++- 3 files changed, 39 insertions(+), 1 deletion(-) diff --git a/src/algebra/order/field/basic.lean b/src/algebra/order/field/basic.lean index f43a92a264d4e..3477d93077b99 100644 --- a/src/algebra/order/field/basic.lean +++ b/src/algebra/order/field/basic.lean @@ -595,6 +595,9 @@ lt_iff_lt_of_le_iff_le $ div_le_iff_of_neg hc lemma lt_div_iff_of_neg' (hc : c < 0) : a < b / c ↔ b < c * a := by rw [mul_comm, lt_div_iff_of_neg hc] +lemma div_le_one_of_ge (h : b ≤ a) (hb : b ≤ 0) : a / b ≤ 1 := +by simpa only [neg_div_neg_eq] using div_le_one_of_le (neg_le_neg h) (neg_nonneg_of_nonpos hb) + /-! ### Bi-implications of inequalities using inversions -/ lemma inv_le_inv_of_neg (ha : a < 0) (hb : b < 0) : a⁻¹ ≤ b⁻¹ ↔ b ≤ a := diff --git a/src/algebra/order/ring/defs.lean b/src/algebra/order/ring/defs.lean index 64e8780473a26..24067d42c45fc 100644 --- a/src/algebra/order/ring/defs.lean +++ b/src/algebra/order/ring/defs.lean @@ -327,6 +327,22 @@ lemma mul_le_mul_of_nonpos_of_nonpos' (hca : c ≤ a) (hdb : d ≤ b) (ha : a a * b ≤ c * d := (mul_le_mul_of_nonpos_left hdb ha).trans $ mul_le_mul_of_nonpos_right hca hd +/-- Variant of `mul_le_of_le_one_left` for `b` non-positive instead of non-negative. -/ +lemma le_mul_of_le_one_left (hb : b ≤ 0) (h : a ≤ 1) : b ≤ a * b := +by simpa only [one_mul] using mul_le_mul_of_nonpos_right h hb + +/-- Variant of `le_mul_of_one_le_left` for `b` non-positive instead of non-negative. -/ +lemma mul_le_of_one_le_left (hb : b ≤ 0) (h : 1 ≤ a) : a * b ≤ b := +by simpa only [one_mul] using mul_le_mul_of_nonpos_right h hb + +/-- Variant of `mul_le_of_le_one_right` for `a` non-positive instead of non-negative. -/ +lemma le_mul_of_le_one_right (ha : a ≤ 0) (h : b ≤ 1) : a ≤ a * b := +by simpa only [mul_one] using mul_le_mul_of_nonpos_left h ha + +/-- Variant of `le_mul_of_one_le_right` for `a` non-positive instead of non-negative. -/ +lemma mul_le_of_one_le_right (ha : a ≤ 0) (h : 1 ≤ b) : a * b ≤ a := +by simpa only [mul_one] using mul_le_mul_of_nonpos_left h ha + section monotone variables [preorder β] {f g : β → α} @@ -568,6 +584,22 @@ by simpa only [mul_neg, neg_lt_neg_iff] using mul_lt_mul_of_pos_right h (neg_pos lemma mul_pos_of_neg_of_neg {a b : α} (ha : a < 0) (hb : b < 0) : 0 < a * b := by simpa only [zero_mul] using mul_lt_mul_of_neg_right ha hb +/-- Variant of `mul_lt_of_lt_one_left` for `b` negative instead of positive. -/ +lemma lt_mul_of_lt_one_left (hb : b < 0) (h : a < 1) : b < a * b := +by simpa only [one_mul] using mul_lt_mul_of_neg_right h hb + +/-- Variant of `lt_mul_of_one_lt_left` for `b` negative instead of positive. -/ +lemma mul_lt_of_one_lt_left (hb : b < 0) (h : 1 < a) : a * b < b := +by simpa only [one_mul] using mul_lt_mul_of_neg_right h hb + +/-- Variant of `mul_lt_of_lt_one_right` for `a` negative instead of positive. -/ +lemma lt_mul_of_lt_one_right (ha : a < 0) (h : b < 1) : a < a * b := +by simpa only [mul_one] using mul_lt_mul_of_neg_left h ha + +/-- Variant of `lt_mul_of_lt_one_right` for `a` negative instead of positive. -/ +lemma mul_lt_of_one_lt_right (ha : a < 0) (h : 1 < b) : a * b < a := +by simpa only [mul_one] using mul_lt_mul_of_neg_left h ha + section monotone variables [preorder β] {f g : β → α} diff --git a/src/algebra/order/ring/lemmas.lean b/src/algebra/order/ring/lemmas.lean index 2b794a53f1b71..cff9f4f845790 100644 --- a/src/algebra/order/ring/lemmas.lean +++ b/src/algebra/order/ring/lemmas.lean @@ -548,7 +548,10 @@ lemma mul_lt_iff_lt_one_left a * b < b ↔ a < 1 := iff.trans (by rw [one_mul]) (mul_lt_mul_right b0) -/-! Lemmas of the form `1 ≤ b → a ≤ a * b`. -/ +/-! Lemmas of the form `1 ≤ b → a ≤ a * b`. + +Variants with `< 0` and `≤ 0` instead of `0 <` and `0 ≤` appear in `algebra/order/ring/defs` (which +imports this file) as they need additional results which are not yet available here. -/ lemma mul_le_of_le_one_left [mul_pos_mono α] (hb : 0 ≤ b) (h : a ≤ 1) : a * b ≤ b := by simpa only [one_mul] using mul_le_mul_of_nonneg_right h hb From 996a85302b992a170cf7336beb4af2d8c3df688c Mon Sep 17 00:00:00 2001 From: MohanadAhmed Date: Thu, 20 Apr 2023 14:40:28 +0000 Subject: [PATCH 0851/1291] feat(linear_algebra/matrix/nonsingular_inverse): interchange of matrix reindexing and inversion (#18812) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This follows the strategy taken in #13827, which gives us all of: * `is_unit (A.submatrix e₁ e₂) ↔ is_unit A` * `(A.submatrix e₁ e₂)⁻¹ = (A⁻¹).submatrix e₂ e₁` * `⅟(A.submatrix e₁ e₂) = (⅟A).submatrix e₂ e₁` * `invertible (A.submatrix e₁ e₂) ≃ invertible A` Co-authored-by: Eric Wieser Co-authored-by: Eric Wieser --- .../matrix/nonsingular_inverse.lean | 72 +++++++++++++++++++ 1 file changed, 72 insertions(+) diff --git a/src/linear_algebra/matrix/nonsingular_inverse.lean b/src/linear_algebra/matrix/nonsingular_inverse.lean index 9493e921678fe..a7231b56c5de2 100644 --- a/src/linear_algebra/matrix/nonsingular_inverse.lean +++ b/src/linear_algebra/matrix/nonsingular_inverse.lean @@ -451,6 +451,8 @@ begin rw [smul_mul, mul_adjugate, units.smul_def, smul_smul, h.coe_inv_mul, one_smul] end +section diagonal + /-- `diagonal v` is invertible if `v` is -/ def diagonal_invertible {α} [non_assoc_semiring α] (v : n → α) [invertible v] : invertible (diagonal v) := @@ -508,6 +510,8 @@ begin rw [ring.inverse_non_unit _ h, pi.zero_def, diagonal_zero, ring.inverse_non_unit _ this] } end +end diagonal + @[simp] lemma inv_inv_inv (A : matrix n n α) : A⁻¹⁻¹⁻¹ = A⁻¹ := begin by_cases h : is_unit A.det, @@ -543,6 +547,74 @@ end by rw [← (A⁻¹).transpose_transpose, vec_mul_transpose, transpose_nonsing_inv, ← det_transpose, Aᵀ.det_smul_inv_mul_vec_eq_cramer _ (is_unit_det_transpose A h)] +/-! ### Inverses of permutated matrices + +Note that the simp-normal form of `matrix.reindex` is `matrix.submatrix`, so we prove most of these +results about only the latter. +-/ + +section submatrix +variables [fintype m] +variables [decidable_eq m] + +/-- `A.submatrix e₁ e₂` is invertible if `A` is -/ +def submatrix_equiv_invertible (A : matrix m m α) (e₁ e₂ : n ≃ m) [invertible A] : + invertible (A.submatrix e₁ e₂) := +invertible_of_right_inverse _ ((⅟A).submatrix e₂ e₁) $ + by rw [matrix.submatrix_mul_equiv, matrix.mul_inv_of_self, submatrix_one_equiv] + +/-- `A` is invertible if `A.submatrix e₁ e₂` is -/ +def invertible_of_submatrix_equiv_invertible (A : matrix m m α) (e₁ e₂ : n ≃ m) + [invertible (A.submatrix e₁ e₂)] : invertible A := +invertible_of_right_inverse _ ((⅟(A.submatrix e₁ e₂)).submatrix e₂.symm e₁.symm) $ begin + have : A = (A.submatrix e₁ e₂).submatrix e₁.symm e₂.symm := by simp, + conv in (_ ⬝ _) { congr, rw this }, + rw [matrix.submatrix_mul_equiv, matrix.mul_inv_of_self, submatrix_one_equiv] +end + +lemma inv_of_submatrix_equiv_eq (A : matrix m m α) (e₁ e₂ : n ≃ m) + [invertible A] [invertible (A.submatrix e₁ e₂)] : + ⅟(A.submatrix e₁ e₂) = (⅟A).submatrix e₂ e₁ := +begin + letI := submatrix_equiv_invertible A e₁ e₂, + haveI := invertible.subsingleton (A.submatrix e₁ e₂), + convert (rfl : ⅟(A.submatrix e₁ e₂) = _), +end + +/-- Together `matrix.submatrix_equiv_invertible` and +`matrix.invertible_of_submatrix_equiv_invertible` form an equivalence, although both sides of the +equiv are subsingleton anyway. -/ +@[simps] +def submatrix_equiv_invertible_equiv_invertible (A : matrix m m α) (e₁ e₂ : n ≃ m) : + invertible (A.submatrix e₁ e₂) ≃ invertible A := +{ to_fun := λ _, by exactI invertible_of_submatrix_equiv_invertible A e₁ e₂, + inv_fun := λ _, by exactI submatrix_equiv_invertible A e₁ e₂, + left_inv := λ _, subsingleton.elim _ _, + right_inv := λ _, subsingleton.elim _ _ } + +/-- When lowered to a prop, `matrix.invertible_of_submatrix_equiv_invertible` forms an `iff`. -/ +@[simp] lemma is_unit_submatrix_equiv {A : matrix m m α} (e₁ e₂ : n ≃ m) : + is_unit (A.submatrix e₁ e₂) ↔ is_unit A := +by simp only [← nonempty_invertible_iff_is_unit, + (submatrix_equiv_invertible_equiv_invertible A _ _).nonempty_congr] + +@[simp] lemma inv_submatrix_equiv (A : matrix m m α) (e₁ e₂ : n ≃ m) : + (A.submatrix e₁ e₂)⁻¹ = (A⁻¹).submatrix e₂ e₁ := +begin + by_cases h : is_unit A, + { casesI h.nonempty_invertible, + letI := submatrix_equiv_invertible A e₁ e₂, + rw [←inv_of_eq_nonsing_inv, ←inv_of_eq_nonsing_inv, inv_of_submatrix_equiv_eq] }, + { have := (is_unit_submatrix_equiv e₁ e₂).not.mpr h, + simp_rw [nonsing_inv_eq_ring_inverse, ring.inverse_non_unit _ h, ring.inverse_non_unit _ this, + submatrix_zero, pi.zero_apply] } +end + +lemma inv_reindex (e₁ e₂ : n ≃ m) (A : matrix n n α) : (reindex e₁ e₂ A)⁻¹ = reindex e₂ e₁ (A⁻¹) := +inv_submatrix_equiv A e₁.symm e₂.symm + +end submatrix + /-! ### More results about determinants -/ section det From 246f6f7989ff86bd07e1b014846f11304f33cf9e Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Thu, 20 Apr 2023 16:22:55 +0000 Subject: [PATCH 0852/1291] feat(algebra/group_with_zero): add lemmas about `f * pi.single i x` (#6418) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Yaël Dillies --- src/algebra/group/pi.lean | 18 +++++++++++++++++- 1 file changed, 17 insertions(+), 1 deletion(-) diff --git a/src/algebra/group/pi.lean b/src/algebra/group/pi.lean index 0039f46c13aba..ce97c2c7e1ac5 100644 --- a/src/algebra/group/pi.lean +++ b/src/algebra/group/pi.lean @@ -22,7 +22,7 @@ universes u v w variables {ι α : Type*} variable {I : Type u} -- The indexing type variable {f : I → Type v} -- The family of types already equipped with instances -variables (x y : Π i, f i) (i : I) +variables (x y : Π i, f i) (i j : I) @[to_additive] lemma set.preimage_one {α β : Type*} [has_one β] (s : set β) [decidable ((1 : β) ∈ s)] : @@ -333,6 +333,22 @@ lemma pi.single_mul [Π i, mul_zero_class $ f i] (i : I) (x y : f i) : single i (x * y) = single i x * single i y := (mul_hom.single f i).map_mul x y +lemma pi.single_mul_left_apply [Π i, mul_zero_class $ f i] (a : f i) : + pi.single i (a * x i) j = pi.single i a j * x j := +(pi.apply_single (λ i, (* x i)) (λ i, zero_mul _) _ _ _).symm + +lemma pi.single_mul_right_apply [Π i, mul_zero_class $ f i] (a : f i) : + pi.single i (x i * a) j = x j * pi.single i a j := +(pi.apply_single (λ i, ((*) (x i))) (λ i, mul_zero _) _ _ _).symm + +lemma pi.single_mul_left [Π i, mul_zero_class $ f i] (a : f i) : + pi.single i (a * x i) = pi.single i a * x := +funext $ λ j, pi.single_mul_left_apply _ _ _ _ + +lemma pi.single_mul_right [Π i, mul_zero_class $ f i] (a : f i) : + pi.single i (x i * a) = x * pi.single i a := +funext $ λ j, pi.single_mul_right_apply _ _ _ _ + /-- The injection into a pi group at different indices commutes. For injections of commuting elements at the same index, see `commute.map` -/ From e8cf0cfec5fcab9baf46dc17d30c5e22048468be Mon Sep 17 00:00:00 2001 From: Hanting Zhang Date: Thu, 20 Apr 2023 22:04:33 +0000 Subject: [PATCH 0853/1291] feat(order/compactly_generated): For any `b`, there exists a set of independent atoms `s` such that `Sup s` is the complement of `b`. (#8475) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This new lemma is carved out of the proof that atomistic lattices are complemented. Also provide directed versions of the interaction between suprema and disjointness. Co-authored-by: Yaël Dillies --- src/order/compactly_generated.lean | 148 ++++++++++++++++++----------- src/order/directed.lean | 2 +- 2 files changed, 95 insertions(+), 55 deletions(-) diff --git a/src/order/compactly_generated.lean b/src/order/compactly_generated.lean index 47ee173833898..bd288f7693805 100644 --- a/src/order/compactly_generated.lean +++ b/src/order/compactly_generated.lean @@ -53,7 +53,11 @@ This is demonstrated by means of the following four lemmas: complete lattice, well-founded, compact -/ -variables {α : Type*} [complete_lattice α] +alias directed_on_range ↔ directed.directed_on_range _ + +attribute [protected] directed.directed_on_range + +variables {ι : Sort*} {α : Type*} [complete_lattice α] {f : ι → α} namespace complete_lattice @@ -333,8 +337,7 @@ theorem le_iff_compact_le_imp {a b : α} : end⟩ /-- This property is sometimes referred to as `α` being upper continuous. -/ -theorem inf_Sup_eq_of_directed_on (h : directed_on (≤) s): - a ⊓ Sup s = ⨆ b ∈ s, a ⊓ b := +theorem directed_on.inf_Sup_eq (h : directed_on (≤) s) : a ⊓ Sup s = ⨆ b ∈ s, a ⊓ b := le_antisymm (begin rw le_iff_compact_le_imp, by_cases hs : s.nonempty, @@ -347,6 +350,32 @@ le_antisymm (begin simp [hs] } end) supr_inf_le_inf_Sup +/-- This property is sometimes referred to as `α` being upper continuous. -/ +protected lemma directed_on.Sup_inf_eq (h : directed_on (≤) s) : Sup s ⊓ a = ⨆ b ∈ s, b ⊓ a := +by simp_rw [@inf_comm _ _ _ a, h.inf_Sup_eq] + +protected lemma directed.inf_supr_eq (h : directed (≤) f) : a ⊓ (⨆ i, f i) = ⨆ i, a ⊓ f i := +by rw [supr, h.directed_on_range.inf_Sup_eq, supr_range] + +protected lemma directed.supr_inf_eq (h : directed (≤) f) : (⨆ i, f i) ⊓ a = ⨆ i, f i ⊓ a := +by rw [supr, h.directed_on_range.Sup_inf_eq, supr_range] + +protected lemma directed_on.disjoint_Sup_right (h : directed_on (≤) s) : + disjoint a (Sup s) ↔ ∀ ⦃b⦄, b ∈ s → disjoint a b := +by simp_rw [disjoint_iff, h.inf_Sup_eq, supr_eq_bot] + +protected lemma directed_on.disjoint_Sup_left (h : directed_on (≤) s) : + disjoint (Sup s) a ↔ ∀ ⦃b⦄, b ∈ s → disjoint b a := +by simp_rw [disjoint_iff, h.Sup_inf_eq, supr_eq_bot] + +protected lemma directed.disjoint_supr_right (h : directed (≤) f) : + disjoint a (⨆ i, f i) ↔ ∀ i, disjoint a (f i) := +by simp_rw [disjoint_iff, h.inf_supr_eq, supr_eq_bot] + +protected lemma directed.disjoint_supr_left (h : directed (≤) f) : + disjoint (⨆ i, f i) a ↔ ∀ i, disjoint (f i) a := +by simp_rw [disjoint_iff, h.supr_inf_eq, supr_eq_bot] + /-- This property is equivalent to `α` being upper continuous. -/ theorem inf_Sup_eq_supr_inf_sup_finset : a ⊓ Sup s = ⨆ (t : finset α) (H : ↑t ⊆ s), a ⊓ (t.sup id) := @@ -460,7 +489,7 @@ instance is_atomic_of_complemented_lattice [complemented_lattice α] : is_atomic exact ⟨a, ha.of_is_atom_coe_Iic, hac.trans hcb⟩ }, end⟩ -/-- See Lemma 5.1, Călugăreanu -/ +/-- See [Lemma 5.1][calugareanu]. -/ @[priority 100] instance is_atomistic_of_complemented_lattice [complemented_lattice α] : is_atomistic α := ⟨λ b, ⟨{a | is_atom a ∧ a ≤ b}, begin @@ -477,59 +506,70 @@ instance is_atomistic_of_complemented_lattice [complemented_lattice α] : is_ato exact le_Sup ⟨ha.of_is_atom_coe_Iic, a.2⟩ } end, λ _, and.left⟩⟩ -/-- See Theorem 6.6, Călugăreanu -/ +/-! +Now we will prove that a compactly generated modular atomistic lattice is a complemented lattice. +Most explicitly, every element is the complement of a supremum of indepedendent atoms. +-/ + +/-- In an atomic lattice, every element `b` has a complement of the form `Sup s`, where each element +of `s` is an atom. See also `complemented_lattice_of_Sup_atoms_eq_top`. -/ +lemma exists_set_independent_is_compl_Sup_atoms (h : Sup {a : α | is_atom a} = ⊤) (b : α) : + ∃ s : set α, complete_lattice.set_independent s ∧ is_compl b (Sup s) ∧ ∀ ⦃a⦄, a ∈ s → is_atom a := +begin + obtain ⟨s, ⟨s_ind, b_inf_Sup_s, s_atoms⟩, s_max⟩ := zorn_subset + {s : set α | complete_lattice.set_independent s ∧ disjoint b (Sup s) ∧ ∀ a ∈ s, is_atom a} + (λ c hc1 hc2, ⟨⋃₀ c, ⟨complete_lattice.independent_sUnion_of_directed hc2.directed_on + (λ s hs, (hc1 hs).1), _, λ a ⟨s, sc, as⟩, (hc1 sc).2.2 a as⟩, λ _, set.subset_sUnion_of_mem⟩), + swap, + { rw [Sup_sUnion, ← Sup_image, directed_on.disjoint_Sup_right], + { rintro _ ⟨s, hs, rfl⟩, + exact (hc1 hs).2.1 }, + { rw directed_on_image, + exact hc2.directed_on.mono (λ s t, Sup_le_Sup) } }, + refine ⟨s, s_ind, ⟨b_inf_Sup_s, _⟩, s_atoms⟩, + rw [codisjoint_iff_le_sup, ←h, Sup_le_iff], + intros a ha, + rw ← inf_eq_left, + refine (ha.le_iff.mp inf_le_left).resolve_left (λ con, ha.1 _), + rw [←con, eq_comm, inf_eq_left], + refine (le_Sup _).trans le_sup_right, + rw ← disjoint_iff at con, + have a_dis_Sup_s : disjoint a (Sup s) := con.mono_right le_sup_right, + rw ← s_max (s ∪ {a}) ⟨λ x hx, _, ⟨_, λ x hx, _⟩⟩ (set.subset_union_left _ _), + { exact set.mem_union_right _ (set.mem_singleton _) }, + { rw [set.mem_union, set.mem_singleton_iff] at hx, + obtain rfl | xa := eq_or_ne x a, + { simp only [set.mem_singleton, set.insert_diff_of_mem, set.union_singleton], + exact con.mono_right ((Sup_le_Sup $ set.diff_subset _ _).trans le_sup_right) }, + { have h : (s ∪ {a}) \ {x} = (s \ {x}) ∪ {a}, + { simp only [set.union_singleton], + rw set.insert_diff_of_not_mem, + rw set.mem_singleton_iff, + exact ne.symm xa }, + rw [h, Sup_union, Sup_singleton], + apply (s_ind (hx.resolve_right xa)).disjoint_sup_right_of_disjoint_sup_left + (a_dis_Sup_s.mono_right _).symm, + rw [← Sup_insert, set.insert_diff_singleton, + set.insert_eq_of_mem (hx.resolve_right xa)] } }, + { rw [Sup_union, Sup_singleton], + exact b_inf_Sup_s.disjoint_sup_right_of_disjoint_sup_left con.symm }, + { rw [set.mem_union, set.mem_singleton_iff] at hx, + obtain hx | rfl := hx, + { exact s_atoms x hx }, + { exact ha } } +end + +lemma exists_set_independent_of_Sup_atoms_eq_top (h : Sup {a : α | is_atom a} = ⊤) : + ∃ s : set α, complete_lattice.set_independent s ∧ Sup s = ⊤ ∧ ∀ ⦃a⦄, a ∈ s → is_atom a := +let ⟨s, s_ind, s_top, s_atoms⟩ := exists_set_independent_is_compl_Sup_atoms h ⊥ in + ⟨s, s_ind, eq_top_of_is_compl_bot s_top.symm, s_atoms⟩ + +/-- See [Theorem 6.6][calugareanu]. -/ theorem complemented_lattice_of_Sup_atoms_eq_top (h : Sup {a : α | is_atom a} = ⊤) : complemented_lattice α := -⟨λ b, begin - obtain ⟨s, ⟨s_ind, b_inf_Sup_s, s_atoms⟩, s_max⟩ := zorn_subset - {s : set α | complete_lattice.set_independent s ∧ b ⊓ Sup s = ⊥ ∧ ∀ a ∈ s, is_atom a} _, - { refine ⟨Sup s, disjoint_iff.mpr b_inf_Sup_s, - codisjoint_iff_le_sup.mpr $ h.symm.trans_le $ Sup_le_iff.2 $ λ a ha, _⟩, - rw ← inf_eq_left, - refine (ha.le_iff.mp inf_le_left).resolve_left (λ con, ha.1 _), - rw [eq_bot_iff, ← con], - refine le_inf (le_refl a) ((le_Sup _).trans le_sup_right), - rw ← disjoint_iff at *, - have a_dis_Sup_s : disjoint a (Sup s) := con.mono_right le_sup_right, - rw ← s_max (s ∪ {a}) ⟨λ x hx, _, ⟨_, λ x hx, _⟩⟩ (set.subset_union_left _ _), - { exact set.mem_union_right _ (set.mem_singleton _) }, - { rw [set.mem_union, set.mem_singleton_iff] at hx, - by_cases xa : x = a, - { simp only [xa, set.mem_singleton, set.insert_diff_of_mem, set.union_singleton], - exact con.mono_right (le_trans (Sup_le_Sup (set.diff_subset s {a})) le_sup_right) }, - { have h : (s ∪ {a}) \ {x} = (s \ {x}) ∪ {a}, - { simp only [set.union_singleton], - rw set.insert_diff_of_not_mem, - rw set.mem_singleton_iff, - exact ne.symm xa }, - rw [h, Sup_union, Sup_singleton], - apply (s_ind (hx.resolve_right xa)).disjoint_sup_right_of_disjoint_sup_left - (a_dis_Sup_s.mono_right _).symm, - rw [← Sup_insert, set.insert_diff_singleton, - set.insert_eq_of_mem (hx.resolve_right xa)] } }, - { rw [Sup_union, Sup_singleton, ← disjoint_iff], - exact b_inf_Sup_s.disjoint_sup_right_of_disjoint_sup_left con.symm }, - { rw [set.mem_union, set.mem_singleton_iff] at hx, - cases hx, - { exact s_atoms x hx }, - { rw hx, - exact ha } } }, - { intros c hc1 hc2, - refine ⟨⋃₀ c, ⟨complete_lattice.independent_sUnion_of_directed hc2.directed_on - (λ s hs, (hc1 hs).1), _, λ a ha, _⟩, λ _, set.subset_sUnion_of_mem⟩, - { rw [Sup_sUnion, ← Sup_image, inf_Sup_eq_of_directed_on, supr_eq_bot], - { intro i, - rw supr_eq_bot, - intro hi, - obtain ⟨x, xc, rfl⟩ := (set.mem_image _ _ _).1 hi, - exact (hc1 xc).2.1 }, - { rw directed_on_image, - refine hc2.directed_on.mono (λ s t, Sup_le_Sup) } }, - { rcases set.mem_sUnion.1 ha with ⟨s, sc, as⟩, - exact (hc1 sc).2.2 a as } } -end⟩ +⟨λ b, let ⟨s, _, s_top, s_atoms⟩ := exists_set_independent_is_compl_Sup_atoms h b in ⟨Sup s, s_top⟩⟩ -/-- See Theorem 6.6, Călugăreanu -/ +/-- See [Theorem 6.6][calugareanu]. -/ theorem complemented_lattice_of_is_atomistic [is_atomistic α] : complemented_lattice α := complemented_lattice_of_Sup_atoms_eq_top Sup_atoms_eq_top diff --git a/src/order/directed.lean b/src/order/directed.lean index 15957a24b5695..b5c5f44c6f8fc 100644 --- a/src/order/directed.lean +++ b/src/order/directed.lean @@ -52,7 +52,7 @@ by simp [directed, directed_on]; refine ball_congr (λ x hx, by simp; refl) alias directed_on_iff_directed ↔ directed_on.directed_coe _ -theorem directed_on_range {f : β → α} : +theorem directed_on_range {f : ι → α} : directed r f ↔ directed_on r (set.range f) := by simp_rw [directed, directed_on, set.forall_range_iff, set.exists_range_iff] From 10ee941346c27bdb5e87bb3535100c0b1f08ac41 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 21 Apr 2023 09:08:53 +0000 Subject: [PATCH 0854/1291] fix(group_theory/monoid_localization): fix timeout (#18846) Splitting this into two definitions seems to make both much faster --- src/group_theory/monoid_localization.lean | 15 ++++++++++----- 1 file changed, 10 insertions(+), 5 deletions(-) diff --git a/src/group_theory/monoid_localization.lean b/src/group_theory/monoid_localization.lean index 2c763e0eefe69..3890cdfbe8f28 100644 --- a/src/group_theory/monoid_localization.lean +++ b/src/group_theory/monoid_localization.lean @@ -1503,8 +1503,11 @@ variables [ordered_cancel_comm_monoid α] {s : submonoid α} {a₁ b₁ : α} {a @[to_additive] lemma mk_le_mk : mk a₁ a₂ ≤ mk b₁ b₂ ↔ ↑b₂ * a₁ ≤ a₂ * b₁ := iff.rfl @[to_additive] lemma mk_lt_mk : mk a₁ a₂ < mk b₁ b₂ ↔ ↑b₂ * a₁ < a₂ * b₁ := iff.rfl -@[to_additive] instance : ordered_cancel_comm_monoid (localization s) := -{ le_refl := λ a, localization.induction_on a $ λ a, le_rfl, +-- declaring this separately to the instance below makes things faster +@[to_additive] instance : partial_order (localization s) := +{ le := (≤), + lt := (<), + le_refl := λ a, localization.induction_on a $ λ a, le_rfl, le_trans := λ a b c, localization.induction_on₃ a b c $ λ a b c hab hbc, begin simp only [mk_le_mk] at ⊢ hab hbc, refine le_of_mul_le_mul_left' _, @@ -1520,8 +1523,10 @@ variables [ordered_cancel_comm_monoid α] {s : submonoid α} {a₁ b₁ : α} {a exact λ hab hba, ⟨1, by rw [hab.antisymm hba]⟩, all_goals { intros, refl }, end, - lt_iff_le_not_le := λ a b, localization.induction_on₂ a b $ λ a b, lt_iff_le_not_le, - mul_le_mul_left := λ a b, localization.induction_on₂ a b $ λ a b hab c, + lt_iff_le_not_le := λ a b, localization.induction_on₂ a b $ λ a b, lt_iff_le_not_le } + +@[to_additive] instance : ordered_cancel_comm_monoid (localization s) := +{ mul_le_mul_left := λ a b, localization.induction_on₂ a b $ λ a b hab c, localization.induction_on c $ λ c, begin simp only [mk_mul, mk_le_mk, submonoid.coe_mul, mul_mul_mul_comm _ _ c.1] at ⊢ hab, exact mul_le_mul_left' hab _, @@ -1530,7 +1535,7 @@ variables [ordered_cancel_comm_monoid α] {s : submonoid α} {a₁ b₁ : α} {a simp only [mk_mul, mk_le_mk, submonoid.coe_mul, mul_mul_mul_comm _ _ a.1] at ⊢ hab, exact le_of_mul_le_mul_left' hab, end, - ..localization.comm_monoid _, ..localization.has_le, ..localization.has_lt } + ..localization.comm_monoid s, ..localization.partial_order } @[to_additive] instance decidable_le [decidable_rel ((≤) : α → α → Prop)] : decidable_rel ((≤) : localization s → localization s → Prop) := From 90df25ded755a2cf9651ea850d1abe429b1e4eb1 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 21 Apr 2023 11:13:01 +0000 Subject: [PATCH 0855/1291] feat(order/monotone/basic): `function.update` is monotone (#18841) and `pi.single` too --- src/algebra/group/pi.lean | 11 +++++++++++ src/order/basic.lean | 6 ++++++ src/order/monotone/basic.lean | 10 ++++++++-- 3 files changed, 25 insertions(+), 2 deletions(-) diff --git a/src/algebra/group/pi.lean b/src/algebra/group/pi.lean index ce97c2c7e1ac5..d64049b2ec257 100644 --- a/src/algebra/group/pi.lean +++ b/src/algebra/group/pi.lean @@ -485,3 +485,14 @@ noncomputable def function.extend_by_one.hom [mul_one_class R] : (ι → R) →* map_mul' := λ f g, by { simpa using function.extend_mul s f g 1 1 } } end extend + +namespace pi +variables [decidable_eq I] [Π i, preorder (f i)] [Π i, has_one (f i)] + +@[to_additive] lemma mul_single_mono : monotone (pi.mul_single i : f i → Π i, f i) := +function.update_mono + +@[to_additive] lemma mul_single_strict_mono : strict_mono (pi.mul_single i : f i → Π i, f i) := +function.update_strict_mono + +end pi diff --git a/src/order/basic.lean b/src/order/basic.lean index f868ff486d619..da3341b7806f3 100644 --- a/src/order/basic.lean +++ b/src/order/basic.lean @@ -593,6 +593,12 @@ lemma update_le_update_iff : function.update x i a ≤ function.update y i b ↔ a ≤ b ∧ ∀ j ≠ i, x j ≤ y j := by simp [update_le_iff] {contextual := tt} +@[simp] lemma update_le_update_iff' : update x i a ≤ update x i b ↔ a ≤ b := +by simp [update_le_update_iff] + +@[simp] lemma update_lt_update_iff : update x i a < update x i b ↔ a < b := +lt_iff_lt_of_le_iff_le' update_le_update_iff' update_le_update_iff' + @[simp] lemma le_update_self_iff : x ≤ update x i a ↔ x i ≤ a := by simp [le_update_iff] @[simp] lemma update_le_self_iff : update x i a ≤ x ↔ a ≤ x i := by simp [update_le_iff] @[simp] lemma lt_update_self_iff : x < update x i a ↔ x i < a := by simp [lt_iff_le_not_le] diff --git a/src/order/monotone/basic.lean b/src/order/monotone/basic.lean index e107e23328433..84433c7deff17 100644 --- a/src/order/monotone/basic.lean +++ b/src/order/monotone/basic.lean @@ -62,7 +62,8 @@ decreasing, strictly decreasing open function order_dual universes u v w -variables {α : Type u} {β : Type v} {γ : Type w} {δ : Type*} {r : α → α → Prop} +variables {ι : Type*} {α : Type u} {β : Type v} {γ : Type w} {δ : Type*} {π : ι → Type*} + {r : α → α → Prop} section monotone_def variables [preorder α] [preorder β] @@ -899,8 +900,13 @@ lemma strict_anti.prod_map (hf : strict_anti f) (hg : strict_anti g) : strict_an end partial_order +/-! ### Pi types -/ + namespace function -variables [preorder α] +variables [preorder α] [decidable_eq ι] [Π i, preorder (π i)] {f : Π i, π i} {i : ι} + +lemma update_mono : monotone (f.update i) := λ a b, update_le_update_iff'.2 +lemma update_strict_mono : strict_mono (f.update i) := λ a b, update_lt_update_iff.2 lemma const_mono : monotone (const β : α → β → α) := λ a b h i, h lemma const_strict_mono [nonempty β] : strict_mono (const β : α → β → α) := λ a b, const_lt_const.2 From eba5bb3155cab51d80af00e8d7d69fa271b1302b Mon Sep 17 00:00:00 2001 From: l534zhan Date: Fri, 21 Apr 2023 14:26:59 +0000 Subject: [PATCH 0856/1291] feat(data/matrix/basic): miscellaneous defs and lemmas (#8289) miscellaneous defs and lemmas Co-authored-by: l534zhan Co-authored-by: Eric Wieser Co-authored-by: l534zhan <84618936+l534zhan@users.noreply.github.com> --- src/data/matrix/basic.lean | 32 +++++++++++++++++++++++++++++--- src/data/matrix/block.lean | 6 +++++- 2 files changed, 34 insertions(+), 4 deletions(-) diff --git a/src/data/matrix/basic.lean b/src/data/matrix/basic.lean index 44e0d83386a43..d681d732fc588 100644 --- a/src/data/matrix/basic.lean +++ b/src/data/matrix/basic.lean @@ -472,6 +472,14 @@ by simp_rw [dot_product, mul_comm] v ⬝ᵥ w = v ⟨⟩ * w ⟨⟩ := by simp [dot_product] +section mul_one_class +variables [mul_one_class α] [add_comm_monoid α] + +lemma dot_product_one (v : n → α) : v ⬝ᵥ 1 = ∑ i, v i := by simp [(⬝ᵥ)] +lemma one_dot_product (v : n → α) : 1 ⬝ᵥ v = ∑ i, v i := by simp [(⬝ᵥ)] + +end mul_one_class + section non_unital_non_assoc_semiring variables [non_unital_non_assoc_semiring α] (u v w : m → α) (x y : n → α) @@ -533,6 +541,14 @@ by convert finset.sum_eq_single i (λ j _, this j) _ using 1; simp end non_unital_non_assoc_semiring_decidable +section non_assoc_semiring +variables [non_assoc_semiring α] + +@[simp] lemma one_dot_product_one : (1 : n → α) ⬝ᵥ 1 = fintype.card n := +by simp [dot_product, fintype.card] + +end non_assoc_semiring + section non_unital_non_assoc_ring variables [non_unital_non_assoc_ring α] (u v w : m → α) @@ -1305,7 +1321,15 @@ by { rw matrix.mul_assoc, simpa only [mul_apply, dot_product, mul_vec] } end non_unital_semiring section non_assoc_semiring -variables [fintype m] [decidable_eq m] [non_assoc_semiring α] +variables [non_assoc_semiring α] + +lemma mul_vec_one [fintype n] (A : matrix m n α) : mul_vec A 1 = λ i, ∑ j, A i j := +by ext; simp [mul_vec, dot_product] + +lemma vec_one_mul [fintype m] (A : matrix m n α) : vec_mul 1 A = λ j, ∑ i, A i j := +by ext; simp [vec_mul, dot_product] + +variables [fintype m] [fintype n] [decidable_eq m] @[simp] lemma one_mul_vec (v : m → α) : mul_vec 1 v = v := by { ext, rw [←diagonal_one, mul_vec_diagonal, one_mul] } @@ -1816,10 +1840,11 @@ lemma submatrix_vec_mul_equiv [fintype l] [fintype m] [non_unital_non_assoc_semi vec_mul v (M.submatrix e₁ e₂) = vec_mul (v ∘ e₁.symm) M ∘ e₂ := funext $ λ i, eq.symm (comp_equiv_symm_dot_product _ _ _) -lemma mul_submatrix_one [fintype n] [fintype o] [non_assoc_semiring α] [decidable_eq o] (e₁ : n ≃ o) +lemma mul_submatrix_one [fintype n] [finite o] [non_assoc_semiring α] [decidable_eq o] (e₁ : n ≃ o) (e₂ : l → o) (M : matrix m n α) : M ⬝ (1 : matrix o o α).submatrix e₁ e₂ = submatrix M id (e₁.symm ∘ e₂) := begin + casesI nonempty_fintype o, let A := M.submatrix id e₁.symm, have : M = A.submatrix id e₁, { simp only [submatrix_submatrix, function.comp.right_id, submatrix_id_id, @@ -1829,10 +1854,11 @@ begin equiv.symm_comp_self], end -lemma one_submatrix_mul [fintype m] [fintype o] [non_assoc_semiring α] [decidable_eq o] (e₁ : l → o) +lemma one_submatrix_mul [fintype m] [finite o] [non_assoc_semiring α] [decidable_eq o] (e₁ : l → o) (e₂ : m ≃ o) (M : matrix m n α) : ((1 : matrix o o α).submatrix e₁ e₂).mul M = submatrix M (e₂.symm ∘ e₁) id := begin + casesI nonempty_fintype o, let A := M.submatrix e₂.symm id, have : M = A.submatrix e₂ id, { simp only [submatrix_submatrix, function.comp.right_id, submatrix_id_id, diff --git a/src/data/matrix/block.lean b/src/data/matrix/block.lean index 2b48884753454..bc311dd65413f 100644 --- a/src/data/matrix/block.lean +++ b/src/data/matrix/block.lean @@ -27,10 +27,14 @@ import data.matrix.basic variables {l m n o p q : Type*} {m' n' p' : o → Type*} variables {R : Type*} {S : Type*} {α : Type*} {β : Type*} -open_locale matrix +open_locale big_operators matrix namespace matrix +lemma dot_product_block [fintype m] [fintype n] [has_mul α] [add_comm_monoid α] (v w : m ⊕ n → α) : + v ⬝ᵥ w = v ∘ sum.inl ⬝ᵥ w ∘ sum.inl + v ∘ sum.inr ⬝ᵥ w ∘ sum.inr := +fintype.sum_sum_type _ + section block_matrices /-- We can form a single large matrix by flattening smaller 'block' matrices of compatible From e49764de5f8377071189dc4fa347ef5d6bb352b1 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 21 Apr 2023 14:27:00 +0000 Subject: [PATCH 0857/1291] feat(linear_algebra/matrix/nonsingular_inverse): decomposition of block matrices with an invertible corner (#18847) This result was already in a `have` statement of an existing proof; the lemma generalizes the indices slightly to not require that the overall matrix is square. --- .../matrix/nonsingular_inverse.lean | 46 +++++++++++++++---- 1 file changed, 36 insertions(+), 10 deletions(-) diff --git a/src/linear_algebra/matrix/nonsingular_inverse.lean b/src/linear_algebra/matrix/nonsingular_inverse.lean index a7231b56c5de2..1bb03fe96276e 100644 --- a/src/linear_algebra/matrix/nonsingular_inverse.lean +++ b/src/linear_algebra/matrix/nonsingular_inverse.lean @@ -48,7 +48,7 @@ matrix inverse, cramer, cramer's rule, adjugate namespace matrix universes u u' v -variables {m : Type u} {n : Type u'} {α : Type v} +variables {l : Type*} {m : Type u} {n : Type u'} {α : Type v} open_locale matrix big_operators open equiv equiv.perm finset @@ -615,6 +615,39 @@ inv_submatrix_equiv A e₁.symm e₂.symm end submatrix +/-! ### Block matrices -/ + +section block +variables [fintype l] +variables [decidable_eq l] +variables [fintype m] +variables [decidable_eq m] + +/-- LDU decomposition of a block matrix with an invertible top-left corner, using the +Schur complement. -/ +lemma from_blocks_eq_of_invertible₁₁ + (A : matrix m m α) (B : matrix m n α) (C : matrix l m α) (D : matrix l n α) [invertible A] : + from_blocks A B C D = + from_blocks 1 0 (C⬝⅟A) 1 ⬝ from_blocks A 0 0 (D - C⬝(⅟A)⬝B) ⬝ from_blocks 1 (⅟A⬝B) 0 1 := +by simp only [from_blocks_multiply, matrix.mul_zero, matrix.zero_mul, add_zero, zero_add, + matrix.one_mul, matrix.mul_one, matrix.inv_of_mul_self, matrix.mul_inv_of_self_assoc, + matrix.mul_inv_of_mul_self_cancel, matrix.mul_assoc, add_sub_cancel'_right] + +/-- LDU decomposition of a block matrix with an invertible bottom-right corner, using the +Schur complement. -/ +lemma from_blocks_eq_of_invertible₂₂ + (A : matrix l m α) (B : matrix l n α) (C : matrix n m α) (D : matrix n n α) [invertible D] : + from_blocks A B C D = + from_blocks 1 (B⬝⅟D) 0 1 ⬝ from_blocks (A - B⬝⅟D⬝C) 0 0 D ⬝ from_blocks 1 0 (⅟D ⬝ C) 1 := +(matrix.reindex (equiv.sum_comm _ _) (equiv.sum_comm _ _)).injective $ by + simpa [reindex_apply, sum_comm_symm, + ←submatrix_mul_equiv _ _ _ (equiv.sum_comm n m), + ←submatrix_mul_equiv _ _ _ (equiv.sum_comm n l), + sum_comm_apply, from_blocks_submatrix_sum_swap_sum_swap] + using from_blocks_eq_of_invertible₁₁ D C B A + +end block + /-! ### More results about determinants -/ section det @@ -634,15 +667,8 @@ by rw [←h.unit_spec, ←coe_units_inv, det_units_conj'] the Schur complement. -/ lemma det_from_blocks₁₁ (A : matrix m m α) (B : matrix m n α) (C : matrix n m α) (D : matrix n n α) [invertible A] : (matrix.from_blocks A B C D).det = det A * det (D - C ⬝ (⅟A) ⬝ B) := -begin - have : from_blocks A B C D = - from_blocks 1 0 (C ⬝ ⅟A) 1 ⬝ from_blocks A 0 0 (D - C ⬝ (⅟A) ⬝ B) ⬝ from_blocks 1 (⅟A ⬝ B) 0 1, - { simp only [from_blocks_multiply, matrix.mul_zero, matrix.zero_mul, add_zero, zero_add, - matrix.one_mul, matrix.mul_one, matrix.inv_of_mul_self, matrix.mul_inv_of_self_assoc, - matrix.mul_inv_of_mul_self_cancel, matrix.mul_assoc, add_sub_cancel'_right] }, - rw [this, det_mul, det_mul, det_from_blocks_zero₂₁, det_from_blocks_zero₂₁, - det_from_blocks_zero₁₂, det_one, det_one, one_mul, one_mul, mul_one], -end +by rw [from_blocks_eq_of_invertible₁₁, det_mul, det_mul, det_from_blocks_zero₂₁, + det_from_blocks_zero₂₁, det_from_blocks_zero₁₂, det_one, det_one, one_mul, one_mul, mul_one] @[simp] lemma det_from_blocks_one₁₁ (B : matrix m n α) (C : matrix n m α) (D : matrix n n α) : (matrix.from_blocks 1 B C D).det = det (D - C ⬝ B) := From 86d1873c01a723aba6788f0b9051ae3d23b4c1c3 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Fri, 21 Apr 2023 16:29:10 +0000 Subject: [PATCH 0858/1291] chore(*): add mathlib4 synchronization comments (#18845) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.category.Module.products` * `algebra.homology.additive` * `algebra.homology.homotopy` * `algebraic_topology.alternating_face_map_complex` * `algebraic_topology.cech_nerve` * `algebraic_topology.dold_kan.faces` * `algebraic_topology.dold_kan.functor_n` * `algebraic_topology.dold_kan.homotopies` * `algebraic_topology.dold_kan.notations` * `algebraic_topology.dold_kan.p_infty` * `algebraic_topology.dold_kan.projections` * `analysis.normed_space.ball_action` * `category_theory.idempotents.homological_complex` * `category_theory.monad.adjunction` * `category_theory.monad.algebra` * `category_theory.monad.limits` * `category_theory.monad.products` * `category_theory.preadditive.eilenberg_moore` * `category_theory.preadditive.opposite` * `category_theory.sites.limits` * `data.complex.cardinality` * `data.matrix.notation` * `data.matrix.reflection` * `data.mv_polynomial.funext` * `linear_algebra.free_module.pid` * `linear_algebra.matrix.determinant` * `linear_algebra.matrix.mv_polynomial` * `linear_algebra.matrix.polynomial` * `linear_algebra.matrix.reindex` * `logic.equiv.transfer_instance` * `ring_theory.chain_of_divisors` * `ring_theory.localization.submodule` * `topology.metric_space.polish` --- src/algebra/category/Module/products.lean | 3 +++ src/algebra/homology/additive.lean | 3 +++ src/algebra/homology/homotopy.lean | 3 +++ src/algebraic_topology/alternating_face_map_complex.lean | 3 +++ src/algebraic_topology/cech_nerve.lean | 3 +++ src/algebraic_topology/dold_kan/faces.lean | 3 +++ src/algebraic_topology/dold_kan/functor_n.lean | 3 +++ src/algebraic_topology/dold_kan/homotopies.lean | 3 +++ src/algebraic_topology/dold_kan/notations.lean | 3 +++ src/algebraic_topology/dold_kan/p_infty.lean | 3 +++ src/algebraic_topology/dold_kan/projections.lean | 3 +++ src/analysis/normed_space/ball_action.lean | 3 +++ src/category_theory/idempotents/homological_complex.lean | 3 +++ src/category_theory/monad/adjunction.lean | 3 +++ src/category_theory/monad/algebra.lean | 3 +++ src/category_theory/monad/limits.lean | 3 +++ src/category_theory/monad/products.lean | 3 +++ src/category_theory/preadditive/eilenberg_moore.lean | 3 +++ src/category_theory/preadditive/opposite.lean | 3 +++ src/category_theory/sites/limits.lean | 3 +++ src/data/complex/cardinality.lean | 3 +++ src/data/matrix/notation.lean | 3 +++ src/data/matrix/reflection.lean | 3 +++ src/data/mv_polynomial/funext.lean | 3 +++ src/linear_algebra/free_module/pid.lean | 3 +++ src/linear_algebra/matrix/determinant.lean | 3 +++ src/linear_algebra/matrix/mv_polynomial.lean | 3 +++ src/linear_algebra/matrix/polynomial.lean | 3 +++ src/linear_algebra/matrix/reindex.lean | 3 +++ src/logic/equiv/transfer_instance.lean | 3 +++ src/ring_theory/chain_of_divisors.lean | 3 +++ src/ring_theory/localization/submodule.lean | 3 +++ src/topology/metric_space/polish.lean | 3 +++ 33 files changed, 99 insertions(+) diff --git a/src/algebra/category/Module/products.lean b/src/algebra/category/Module/products.lean index 1db567c743049..96b8558b7a776 100644 --- a/src/algebra/category/Module/products.lean +++ b/src/algebra/category/Module/products.lean @@ -8,6 +8,9 @@ import algebra.category.Module.basic /-! # The concrete products in the category of modules are products in the categorical sense. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ open category_theory diff --git a/src/algebra/homology/additive.lean b/src/algebra/homology/additive.lean index 9ecd5b5fab41d..d7df969443b34 100644 --- a/src/algebra/homology/additive.lean +++ b/src/algebra/homology/additive.lean @@ -10,6 +10,9 @@ import category_theory.preadditive.additive_functor /-! # Homology is an additive functor +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + When `V` is preadditive, `homological_complex V c` is also preadditive, and `homology_functor` is additive. diff --git a/src/algebra/homology/homotopy.lean b/src/algebra/homology/homotopy.lean index 0c28c5eb5b6ae..df7241c4395d3 100644 --- a/src/algebra/homology/homotopy.lean +++ b/src/algebra/homology/homotopy.lean @@ -9,6 +9,9 @@ import tactic.abel /-! # Chain homotopies +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define chain homotopies, and prove that homotopic chain maps induce the same map on homology. -/ diff --git a/src/algebraic_topology/alternating_face_map_complex.lean b/src/algebraic_topology/alternating_face_map_complex.lean index b2a46b05bfbea..9556037c8df4a 100644 --- a/src/algebraic_topology/alternating_face_map_complex.lean +++ b/src/algebraic_topology/alternating_face_map_complex.lean @@ -15,6 +15,9 @@ import tactic.equiv_rw # The alternating face map complex of a simplicial object in a preadditive category +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We construct the alternating face map complex, as a functor `alternating_face_map_complex : simplicial_object C ⥤ chain_complex C ℕ` for any preadditive category `C`. For any simplicial object `X` in `C`, diff --git a/src/algebraic_topology/cech_nerve.lean b/src/algebraic_topology/cech_nerve.lean index fd2452124d894..e6024e9ae55ca 100644 --- a/src/algebraic_topology/cech_nerve.lean +++ b/src/algebraic_topology/cech_nerve.lean @@ -13,6 +13,9 @@ import category_theory.arrow # The Čech Nerve +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file provides a definition of the Čech nerve associated to an arrow, provided the base category has the correct wide pullbacks. diff --git a/src/algebraic_topology/dold_kan/faces.lean b/src/algebraic_topology/dold_kan/faces.lean index ee62e3c18ed85..ae2b2f6e8fff0 100644 --- a/src/algebraic_topology/dold_kan/faces.lean +++ b/src/algebraic_topology/dold_kan/faces.lean @@ -11,6 +11,9 @@ import tactic.ring_exp # Study of face maps for the Dold-Kan correspondence +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + TODO (@joelriou) continue adding the various files referenced below In this file, we obtain the technical lemmas that are used in the file diff --git a/src/algebraic_topology/dold_kan/functor_n.lean b/src/algebraic_topology/dold_kan/functor_n.lean index 2f0888cc7c0d4..17a1db8bb19eb 100644 --- a/src/algebraic_topology/dold_kan/functor_n.lean +++ b/src/algebraic_topology/dold_kan/functor_n.lean @@ -10,6 +10,9 @@ import algebraic_topology.dold_kan.p_infty # Construction of functors N for the Dold-Kan correspondence +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + TODO (@joelriou) continue adding the various files referenced below In this file, we construct functors `N₁ : simplicial_object C ⥤ karoubi (chain_complex C ℕ)` diff --git a/src/algebraic_topology/dold_kan/homotopies.lean b/src/algebraic_topology/dold_kan/homotopies.lean index 5bc1fd2472a28..13533ae3f6936 100644 --- a/src/algebraic_topology/dold_kan/homotopies.lean +++ b/src/algebraic_topology/dold_kan/homotopies.lean @@ -11,6 +11,9 @@ import algebraic_topology.dold_kan.notations # Construction of homotopies for the Dold-Kan correspondence +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + TODO (@joelriou) continue adding the various files referenced below (The general strategy of proof of the Dold-Kan correspondence is explained diff --git a/src/algebraic_topology/dold_kan/notations.lean b/src/algebraic_topology/dold_kan/notations.lean index 0bfe1a14e6162..67fda8f17e311 100644 --- a/src/algebraic_topology/dold_kan/notations.lean +++ b/src/algebraic_topology/dold_kan/notations.lean @@ -10,6 +10,9 @@ import algebraic_topology.alternating_face_map_complex # Notations for the Dold-Kan equivalence +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the notation `K[X] : chain_complex C ℕ` for the alternating face map complex of `(X : simplicial_object C)` where `C` is a preadditive category, as well as `N[X]` for the normalized subcomplex in the case `C` is an abelian category. diff --git a/src/algebraic_topology/dold_kan/p_infty.lean b/src/algebraic_topology/dold_kan/p_infty.lean index 9cd1e7881ba99..b95df628099e3 100644 --- a/src/algebraic_topology/dold_kan/p_infty.lean +++ b/src/algebraic_topology/dold_kan/p_infty.lean @@ -12,6 +12,9 @@ import category_theory.idempotents.functor_extension # Construction of the projection `P_infty` for the Dold-Kan correspondence +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + TODO (@joelriou) continue adding the various files referenced below In this file, we construct the projection `P_infty : K[X] ⟶ K[X]` by passing diff --git a/src/algebraic_topology/dold_kan/projections.lean b/src/algebraic_topology/dold_kan/projections.lean index 6838b6958d465..2439974a9359a 100644 --- a/src/algebraic_topology/dold_kan/projections.lean +++ b/src/algebraic_topology/dold_kan/projections.lean @@ -11,6 +11,9 @@ import category_theory.idempotents.basic # Construction of projections for the Dold-Kan correspondence +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + TODO (@joelriou) continue adding the various files referenced below In this file, we construct endomorphisms `P q : K[X] ⟶ K[X]` for all diff --git a/src/analysis/normed_space/ball_action.lean b/src/analysis/normed_space/ball_action.lean index 472cc334c00a8..99295ec561d81 100644 --- a/src/analysis/normed_space/ball_action.lean +++ b/src/analysis/normed_space/ball_action.lean @@ -9,6 +9,9 @@ import analysis.normed_space.basic /-! # Multiplicative actions of/on balls and spheres +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Let `E` be a normed vector space over a normed field `𝕜`. In this file we define the following multiplicative actions. diff --git a/src/category_theory/idempotents/homological_complex.lean b/src/category_theory/idempotents/homological_complex.lean index ebc8a609b0b24..bbfffb7bf85ca 100644 --- a/src/category_theory/idempotents/homological_complex.lean +++ b/src/category_theory/idempotents/homological_complex.lean @@ -10,6 +10,9 @@ import category_theory.idempotents.karoubi /-! # Idempotent completeness and homological complexes +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains simplifications lemmas for categories `karoubi (homological_complex C c)` and the construction of an equivalence of categories `karoubi (homological_complex C c) ≌ homological_complex (karoubi C) c`. diff --git a/src/category_theory/monad/adjunction.lean b/src/category_theory/monad/adjunction.lean index 8f2ee003612a8..e492d3f7a6923 100644 --- a/src/category_theory/monad/adjunction.lean +++ b/src/category_theory/monad/adjunction.lean @@ -9,6 +9,9 @@ import category_theory.monad.algebra /-! # Adjunctions and monads +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We develop the basic relationship between adjunctions and monads. Given an adjunction `h : L ⊣ R`, we have `h.to_monad : monad C` and `h.to_comonad : comonad D`. diff --git a/src/category_theory/monad/algebra.lean b/src/category_theory/monad/algebra.lean index 198f5f83f93e2..4f034f3323988 100644 --- a/src/category_theory/monad/algebra.lean +++ b/src/category_theory/monad/algebra.lean @@ -10,6 +10,9 @@ import category_theory.functor.epi_mono /-! # Eilenberg-Moore (co)algebras for a (co)monad +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines Eilenberg-Moore (co)algebras for a (co)monad, and provides the category instance for them. diff --git a/src/category_theory/monad/limits.lean b/src/category_theory/monad/limits.lean index 37d516fef5dc9..31bac6c46e2d7 100644 --- a/src/category_theory/monad/limits.lean +++ b/src/category_theory/monad/limits.lean @@ -10,6 +10,9 @@ import category_theory.limits.shapes.terminal /-! # Limits and colimits in the category of algebras +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file shows that the forgetful functor `forget T : algebra T ⥤ C` for a monad `T : C ⥤ C` creates limits and creates any colimits which `T` preserves. This is used to show that `algebra T` has any limits which `C` has, and any colimits which `C` has diff --git a/src/category_theory/monad/products.lean b/src/category_theory/monad/products.lean index fea2f76ee9680..2062ee02cf3be 100644 --- a/src/category_theory/monad/products.lean +++ b/src/category_theory/monad/products.lean @@ -10,6 +10,9 @@ import category_theory.limits.shapes.binary_products /-! # Algebras for the coproduct monad +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The functor `Y ↦ X ⨿ Y` forms a monad, whose category of monads is equivalent to the under category of `X`. Similarly, `Y ↦ X ⨯ Y` forms a comonad, whose category of comonads is equivalent to the over category of `X`. diff --git a/src/category_theory/preadditive/eilenberg_moore.lean b/src/category_theory/preadditive/eilenberg_moore.lean index a4f9535fda0a4..fdb3f3332c170 100644 --- a/src/category_theory/preadditive/eilenberg_moore.lean +++ b/src/category_theory/preadditive/eilenberg_moore.lean @@ -11,6 +11,9 @@ import category_theory.preadditive.additive_functor /-! # Preadditive structure on algebras over a monad +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + If `C` is a preadditive categories and `T` is an additive monad on `C` then `algebra T` is also preadditive. Dually, if `U` is an additive comonad on `C` then `coalgebra U` is preadditive as well. diff --git a/src/category_theory/preadditive/opposite.lean b/src/category_theory/preadditive/opposite.lean index 1dd70b5fc7952..f846ccb1ba59d 100644 --- a/src/category_theory/preadditive/opposite.lean +++ b/src/category_theory/preadditive/opposite.lean @@ -9,6 +9,9 @@ import logic.equiv.transfer_instance /-! # If `C` is preadditive, `Cᵒᵖ` has a natural preadditive structure. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + -/ open opposite diff --git a/src/category_theory/sites/limits.lean b/src/category_theory/sites/limits.lean index 49efac80318e1..8fcf057055190 100644 --- a/src/category_theory/sites/limits.lean +++ b/src/category_theory/sites/limits.lean @@ -10,6 +10,9 @@ import category_theory.sites.sheafification # Limits and colimits of sheaves +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Limits We prove that the forgetful functor from `Sheaf J D` to presheaves creates limits. diff --git a/src/data/complex/cardinality.lean b/src/data/complex/cardinality.lean index 1bbabbe9cde5e..7b1ff847899bc 100644 --- a/src/data/complex/cardinality.lean +++ b/src/data/complex/cardinality.lean @@ -10,6 +10,9 @@ import data.real.cardinality /-! # The cardinality of the complex numbers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file shows that the complex numbers have cardinality continuum, i.e. `#ℂ = 𝔠`. -/ diff --git a/src/data/matrix/notation.lean b/src/data/matrix/notation.lean index dd35474e61c41..8ff6eefcde35a 100644 --- a/src/data/matrix/notation.lean +++ b/src/data/matrix/notation.lean @@ -11,6 +11,9 @@ import algebra.big_operators.fin /-! # Matrix and vector notation +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file includes `simp` lemmas for applying operations in `data.matrix.basic` to values built out of the matrix notation `![a, b] = vec_cons a (vec_cons b vec_empty)` defined in `data.fin.vec_notation`. diff --git a/src/data/matrix/reflection.lean b/src/data/matrix/reflection.lean index dd9c6d06a48ec..690bed1a0d04a 100644 --- a/src/data/matrix/reflection.lean +++ b/src/data/matrix/reflection.lean @@ -10,6 +10,9 @@ import data.fin.tuple.reflection /-! # Lemmas for concrete matrices `matrix (fin m) (fin n) α` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains alternative definitions of common operators on matrices that expand definitionally to the expected expression when evaluated on `!![]` notation. diff --git a/src/data/mv_polynomial/funext.lean b/src/data/mv_polynomial/funext.lean index ed08c70b9336c..455d8d8ae6d8a 100644 --- a/src/data/mv_polynomial/funext.lean +++ b/src/data/mv_polynomial/funext.lean @@ -10,6 +10,9 @@ import ring_theory.polynomial.basic /-! ## Function extensionality for multivariate polynomials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we show that two multivariate polynomials over an infinite integral domain are equal if they are equal upon evaluating them on an arbitrary assignment of the variables. diff --git a/src/linear_algebra/free_module/pid.lean b/src/linear_algebra/free_module/pid.lean index 9f54c0e2859e8..af76def8e465f 100644 --- a/src/linear_algebra/free_module/pid.lean +++ b/src/linear_algebra/free_module/pid.lean @@ -11,6 +11,9 @@ import ring_theory.finiteness /-! # Free modules over PID +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A free `R`-module `M` is a module with a basis over `R`, equivalently it is an `R`-module linearly equivalent to `ι →₀ R` for some `ι`. diff --git a/src/linear_algebra/matrix/determinant.lean b/src/linear_algebra/matrix/determinant.lean index 96afa0d0acd05..fa527c57a6b8c 100644 --- a/src/linear_algebra/matrix/determinant.lean +++ b/src/linear_algebra/matrix/determinant.lean @@ -17,6 +17,9 @@ import linear_algebra.pi /-! # Determinant of a matrix +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the determinant of a matrix, `matrix.det`, and its essential properties. ## Main definitions diff --git a/src/linear_algebra/matrix/mv_polynomial.lean b/src/linear_algebra/matrix/mv_polynomial.lean index 8c02733b22a3b..64a648d3bd5de 100644 --- a/src/linear_algebra/matrix/mv_polynomial.lean +++ b/src/linear_algebra/matrix/mv_polynomial.lean @@ -10,6 +10,9 @@ import data.mv_polynomial.comm_ring /-! # Matrices of multivariate polynomials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we prove results about matrices over an mv_polynomial ring. In particular, we provide `matrix.mv_polynomial_X` which associates every entry of a matrix with a unique variable. diff --git a/src/linear_algebra/matrix/polynomial.lean b/src/linear_algebra/matrix/polynomial.lean index 0ee53e6f73750..4045bbe16a2cb 100644 --- a/src/linear_algebra/matrix/polynomial.lean +++ b/src/linear_algebra/matrix/polynomial.lean @@ -10,6 +10,9 @@ import linear_algebra.matrix.determinant /-! # Matrices of polynomials and polynomials of matrices +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we prove results about matrices over a polynomial ring. In particular, we give results about the polynomial given by `det (t * I + A)`. diff --git a/src/linear_algebra/matrix/reindex.lean b/src/linear_algebra/matrix/reindex.lean index 0df5c98bf91af..10513e866f552 100644 --- a/src/linear_algebra/matrix/reindex.lean +++ b/src/linear_algebra/matrix/reindex.lean @@ -9,6 +9,9 @@ import linear_algebra.matrix.determinant /-! # Changing the index type of a matrix +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file concerns the map `matrix.reindex`, mapping a `m` by `n` matrix to an `m'` by `n'` matrix, as long as `m ≃ m'` and `n ≃ n'`. diff --git a/src/logic/equiv/transfer_instance.lean b/src/logic/equiv/transfer_instance.lean index 1db07027c0132..d20b4270e3382 100644 --- a/src/logic/equiv/transfer_instance.lean +++ b/src/logic/equiv/transfer_instance.lean @@ -10,6 +10,9 @@ import logic.equiv.defs /-! # Transfer algebraic structures across `equiv`s +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove theorems of the following form: if `β` has a group structure and `α ≃ β` then `α` has a group structure, and similarly for monoids, semigroups, rings, integral domains, fields and diff --git a/src/ring_theory/chain_of_divisors.lean b/src/ring_theory/chain_of_divisors.lean index 8a729133e178c..9c69310477c7f 100644 --- a/src/ring_theory/chain_of_divisors.lean +++ b/src/ring_theory/chain_of_divisors.lean @@ -12,6 +12,9 @@ import algebra.gcd_monoid.basic # Chains of divisors +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The results in this file show that in the monoid `associates M` of a `unique_factorization_monoid` `M`, an element `a` is an n-th prime power iff its set of divisors is a strictly increasing chain of length `n + 1`, meaning that we can find a strictly increasing bijection between `fin (n + 1)` diff --git a/src/ring_theory/localization/submodule.lean b/src/ring_theory/localization/submodule.lean index ffe41ebd8441f..ef2f6f717de87 100644 --- a/src/ring_theory/localization/submodule.lean +++ b/src/ring_theory/localization/submodule.lean @@ -10,6 +10,9 @@ import ring_theory.principal_ideal_domain /-! # Submodules in localizations of commutative rings +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Implementation notes See `src/ring_theory/localization/basic.lean` for a design overview. diff --git a/src/topology/metric_space/polish.lean b/src/topology/metric_space/polish.lean index cfcf50428ebda..e4cc253a4a45e 100644 --- a/src/topology/metric_space/polish.lean +++ b/src/topology/metric_space/polish.lean @@ -11,6 +11,9 @@ import analysis.normed.field.basic /-! # Polish spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A topological space is Polish if its topology is second-countable and there exists a compatible complete metric. This is the class of spaces that is well-behaved with respect to measure theory. In this file, we establish the basic properties of Polish spaces. From 52932b3a083d4142e78a15dc928084a22fea9ba0 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 21 Apr 2023 19:31:18 +0000 Subject: [PATCH 0859/1291] chore(measure_theory/function/ess_sup): Generalise (#18669) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit A handful of lemmas hold for bounded filters in conditionally complete lattices, rather than just filter in complete lattices (which are automatically bounded). Also prove that `μ {y | x < f y} = 0` when `x` is greater than the essential supremum of `f`, and dually. Co-authored-by: sgouezel --- src/measure_theory/function/ess_sup.lean | 72 ++++++++++++------- src/measure_theory/function/lp_space.lean | 5 ++ src/order/filter/ennreal.lean | 32 +-------- src/topology/algebra/order/liminf_limsup.lean | 40 ++++++++++- 4 files changed, 95 insertions(+), 54 deletions(-) diff --git a/src/measure_theory/function/ess_sup.lean b/src/measure_theory/function/ess_sup.lean index 1f8e498fb7586..23ecbb3325864 100644 --- a/src/measure_theory/function/ess_sup.lean +++ b/src/measure_theory/function/ess_sup.lean @@ -26,8 +26,8 @@ sense). We do not define that quantity here, which is simply the supremum of a m * `ess_inf f μ := μ.ae.liminf f` -/ -open measure_theory filter topological_space -open_locale ennreal measure_theory +open measure_theory filter set topological_space +open_locale ennreal measure_theory nnreal variables {α β : Type*} {m : measurable_space α} {μ ν : measure α} @@ -48,19 +48,53 @@ limsup_congr hfg lemma ess_inf_congr_ae {f g : α → β} (hfg : f =ᵐ[μ] g) : ess_inf f μ = ess_inf g μ := @ess_sup_congr_ae α βᵒᵈ _ _ _ _ _ hfg +@[simp] lemma ess_sup_const' [μ.ae.ne_bot] (c : β) : ess_sup (λ x : α, c) μ = c := limsup_const _ +@[simp] lemma ess_inf_const' [μ.ae.ne_bot] (c : β) : ess_inf (λ x : α, c) μ = c := liminf_const _ + +lemma ess_sup_const (c : β) (hμ : μ ≠ 0) : ess_sup (λ x : α, c) μ = c := +by { rw ←ae_ne_bot at hμ, exactI ess_sup_const' _ } + +lemma ess_inf_const (c : β) (hμ : μ ≠ 0) : ess_inf (λ x : α, c) μ = c := +by { rw ←ae_ne_bot at hμ, exactI ess_inf_const' _ } + end conditionally_complete_lattice section conditionally_complete_linear_order -variable [conditionally_complete_linear_order β] +variables [conditionally_complete_linear_order β] {x : β} {f : α → β} lemma ess_sup_eq_Inf {m : measurable_space α} (μ : measure α) (f : α → β) : ess_sup f μ = Inf {a | μ {x | a < f x} = 0} := -begin - dsimp [ess_sup, limsup, Limsup], - congr, - ext a, - simp [eventually_map, ae_iff], -end +by { dsimp [ess_sup, limsup, Limsup], simp only [ae_iff, not_le] } + +lemma ess_inf_eq_Sup {m : measurable_space α} (μ : measure α) (f : α → β) : + ess_inf f μ = Sup {a | μ {x | f x < a} = 0} := +by { dsimp [ess_inf, liminf, Liminf], simp only [ae_iff, not_le] } + +lemma ae_lt_of_ess_sup_lt (hx : ess_sup f μ < x) + (hf : is_bounded_under (≤) μ.ae f . is_bounded_default) : ∀ᵐ y ∂μ, f y < x := +eventually_lt_of_limsup_lt hx hf + +lemma ae_lt_of_lt_ess_inf (hx : x < ess_inf f μ) + (hf : is_bounded_under (≥) μ.ae f . is_bounded_default) : ∀ᵐ y ∂μ, x < f y := +eventually_lt_of_lt_liminf hx hf + +variables [topological_space β] [first_countable_topology β] [order_topology β] + +lemma ae_le_ess_sup (hf : is_bounded_under (≤) μ.ae f . is_bounded_default) : + ∀ᵐ y ∂μ, f y ≤ ess_sup f μ := +eventually_le_limsup hf + +lemma ae_ess_inf_le (hf : is_bounded_under (≥) μ.ae f . is_bounded_default) : + ∀ᵐ y ∂μ, ess_inf f μ ≤ f y := +eventually_liminf_le hf + +lemma meas_ess_sup_lt (hf : is_bounded_under (≤) μ.ae f . is_bounded_default) : + μ {y | ess_sup f μ < f y} = 0 := +by { simp_rw ←not_le, exact ae_le_ess_sup hf } + +lemma meas_lt_ess_inf (hf : is_bounded_under (≥) μ.ae f . is_bounded_default) : + μ {y | f y < ess_inf f μ} = 0 := +by { simp_rw ←not_le, exact ae_ess_inf_le hf } end conditionally_complete_linear_order @@ -81,12 +115,6 @@ limsup_le_limsup hfg lemma ess_inf_mono_ae {f g : α → β} (hfg : f ≤ᵐ[μ] g) : ess_inf f μ ≤ ess_inf g μ := liminf_le_liminf hfg -lemma ess_sup_const (c : β) (hμ : μ ≠ 0) : ess_sup (λ x : α, c) μ = c := -begin - haveI hμ_ne_bot : μ.ae.ne_bot, { rwa [ne_bot_iff, ne.def, ae_eq_bot] }, - exact limsup_const c, -end - lemma ess_sup_le_of_ae_le {f : α → β} (c : β) (hf : f ≤ᵐ[μ] (λ _, c)) : ess_sup f μ ≤ c := begin refine (ess_sup_mono_ae hf).trans _, @@ -95,9 +123,6 @@ begin { rwa ess_sup_const, }, end -lemma ess_inf_const (c : β) (hμ : μ ≠ 0) : ess_inf (λ x : α, c) μ = c := -@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 α βᵒᵈ _ _ _ _ c hf @@ -203,12 +228,6 @@ end complete_lattice section complete_linear_order variable [complete_linear_order β] -lemma ae_lt_of_ess_sup_lt {f : α → β} {x : β} (hf : ess_sup f μ < x) : ∀ᵐ y ∂μ, f y < 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 α βᵒᵈ _ _ _ _ _ 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) : ess_sup (s.indicator f) μ = ess_sup f (μ.restrict s) := @@ -262,4 +281,9 @@ lemma ess_sup_liminf_le {ι} [countable ι] [linear_order ι] (f : ι → α → ess_sup (λ x, at_top.liminf (λ n, f n x)) μ ≤ at_top.liminf (λ n, ess_sup (λ x, f n x) μ) := by { simp_rw ess_sup, exact ennreal.limsup_liminf_le_liminf_limsup (λ a b, f b a), } +lemma coe_ess_sup {f : α → ℝ≥0} (hf : is_bounded_under (≤) μ.ae f) : + (↑(ess_sup f μ) : ℝ≥0∞) = ess_sup (λ x, f x) μ := +(ennreal.coe_Inf $ by exact hf).trans $ eq_of_forall_le_iff $ λ r, + by simp [ess_sup, limsup, Limsup, eventually_map, ennreal.forall_ennreal] + end ennreal diff --git a/src/measure_theory/function/lp_space.lean b/src/measure_theory/function/lp_space.lean index 851f9dc88bdf9..0ff13d618a5e1 100644 --- a/src/measure_theory/function/lp_space.lean +++ b/src/measure_theory/function/lp_space.lean @@ -841,6 +841,11 @@ begin exact ((ennreal.add_lt_top).2 ⟨hf.2, hg.2⟩).ne, end +lemma ae_le_snorm_ess_sup {f : α → F} : ∀ᵐ y ∂μ, ↑‖f y‖₊ ≤ snorm_ess_sup f μ := ae_le_ess_sup + +lemma meas_snorm_ess_sup_lt {f : α → F} : μ {y | snorm_ess_sup f μ < ‖f y‖₊} = 0 := +meas_ess_sup_lt + section map_measure variables {β : Type*} {mβ : measurable_space β} {f : α → β} {g : β → E} diff --git a/src/order/filter/ennreal.lean b/src/order/filter/ennreal.lean index 08033fb22005a..7a9f03cf5b166 100644 --- a/src/order/filter/ennreal.lean +++ b/src/order/filter/ennreal.lean @@ -3,10 +3,7 @@ Copyright (c) 2021 Rémy Degenne. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ - -import data.real.ennreal -import order.filter.countable_Inter -import order.liminf_limsup +import topology.instances.ennreal /-! # Order properties of extended non-negative reals @@ -25,34 +22,11 @@ variables {α : Type*} {f : filter α} lemma eventually_le_limsup [countable_Inter_filter f] (u : α → ℝ≥0∞) : ∀ᶠ y in f, u y ≤ f.limsup u := -begin - by_cases hx_top : f.limsup u = ⊤, - { simp_rw hx_top, - exact eventually_of_forall (λ a, le_top), }, - have h_forall_le : ∀ᶠ y in f, ∀ n : ℕ, u y < f.limsup u + (1:ℝ≥0∞)/n, - { rw eventually_countable_forall, - refine λ n, eventually_lt_of_limsup_lt _, - nth_rewrite 0 ←add_zero (f.limsup u), - exact (ennreal.add_lt_add_iff_left hx_top).mpr (by simp), }, - refine h_forall_le.mono (λ y hy, le_of_forall_pos_le_add (λ r hr_pos hx_top, _)), - have hr_ne_zero : (r : ℝ≥0∞) ≠ 0, - { rw [ne.def, coe_eq_zero], - exact (ne_of_lt hr_pos).symm, }, - cases (exists_inv_nat_lt hr_ne_zero) with i hi, - rw inv_eq_one_div at hi, - exact (hy i).le.trans (add_le_add_left hi.le (f.limsup u)), -end +eventually_le_limsup lemma limsup_eq_zero_iff [countable_Inter_filter f] {u : α → ℝ≥0∞} : f.limsup u = 0 ↔ u =ᶠ[f] 0 := -begin - split; intro h, - { have hu_zero := eventually_le.trans (eventually_le_limsup u) - (eventually_of_forall (λ _, le_of_eq h)), - exact hu_zero.mono (λ x hx, le_antisymm hx (zero_le _)), }, - { rw limsup_congr h, - simp_rw [pi.zero_apply, ←ennreal.bot_eq_zero, limsup_const_bot] }, -end +limsup_eq_bot lemma limsup_const_mul_of_ne_top {u : α → ℝ≥0∞} {a : ℝ≥0∞} (ha_top : a ≠ ⊤) : f.limsup (λ (x : α), a * (u x)) = a * f.limsup u := diff --git a/src/topology/algebra/order/liminf_limsup.lean b/src/topology/algebra/order/liminf_limsup.lean index 8d1cbcb99ac54..4e819a7cdbfb6 100644 --- a/src/topology/algebra/order/liminf_limsup.lean +++ b/src/topology/algebra/order/liminf_limsup.lean @@ -8,6 +8,7 @@ import algebra.big_operators.order import algebra.indicator_function import order.liminf_limsup import order.filter.archimedean +import order.filter.countable_Inter import topology.order.basic /-! @@ -17,7 +18,7 @@ import topology.order.basic > Any changes to this file require a corresponding PR to mathlib4. -/ -open filter +open filter topological_space open_locale topology classical universes u v @@ -215,8 +216,45 @@ begin exact H a as b bs ab ⟨A, B⟩, end +variables [first_countable_topology α] {f : filter β} [countable_Inter_filter f] {u : β → α} + +lemma eventually_le_limsup (hf : is_bounded_under (≤) f u . is_bounded_default) : + ∀ᶠ b in f, u b ≤ f.limsup u := +begin + obtain ha | ha := is_top_or_exists_gt (f.limsup u), + { exact eventually_of_forall (λ _, ha _) }, + by_cases H : is_glb (set.Ioi (f.limsup u)) (f.limsup u), + { obtain ⟨u, -, -, hua, hu⟩ := H.exists_seq_antitone_tendsto ha, + have := λ n, eventually_lt_of_limsup_lt (hu n) hf, + exact (eventually_countable_forall.2 this).mono + (λ b hb, ge_of_tendsto hua $ eventually_of_forall $ λ n, (hb _).le) }, + { obtain ⟨x, hx, xa⟩ : ∃ x, (∀ ⦃b⦄, f.limsup u < b → x ≤ b) ∧ f.limsup u < x, + { simp only [is_glb, is_greatest, lower_bounds, upper_bounds, set.mem_Ioi, set.mem_set_of_eq, + not_and, not_forall, not_le, exists_prop] at H, + exact H (λ x hx, le_of_lt hx) }, + filter_upwards [eventually_lt_of_limsup_lt xa hf] with y hy, + contrapose! hy, + exact hx hy } +end + +lemma eventually_liminf_le (hf : is_bounded_under (≥) f u . is_bounded_default) : + ∀ᶠ b in f, f.liminf u ≤ u b := +@eventually_le_limsup αᵒᵈ _ _ _ _ _ _ _ _ hf + end conditionally_complete_linear_order +section complete_linear_order +variables [complete_linear_order α] [topological_space α] [first_countable_topology α] + [order_topology α] {f : filter β} [countable_Inter_filter f] {u : β → α} + +@[simp] lemma limsup_eq_bot : f.limsup u = ⊥ ↔ u =ᶠ[f] ⊥ := +⟨λ h, (eventually_le.trans eventually_le_limsup $ eventually_of_forall $ λ _, h.le).mono $ λ x hx, + le_antisymm hx bot_le, λ h, by { rw limsup_congr h, exact limsup_const_bot }⟩ + +@[simp] lemma liminf_eq_top : f.liminf u = ⊤ ↔ u =ᶠ[f] ⊤ := @limsup_eq_bot αᵒᵈ _ _ _ _ _ _ _ _ + +end complete_linear_order + end liminf_limsup section monotone From d4817f8867c368d6c5571f7379b3888aaec1d95a Mon Sep 17 00:00:00 2001 From: sgouezel Date: Sat, 22 Apr 2023 06:00:37 +0000 Subject: [PATCH 0860/1291] feat(measure_theory/integral/integral_eq_improper): fundamental theorem of calculus on (a, +\infty) (#18844) It became apparent in several applications that we are missing API here. As an illustration, two existing proofs in the library are shortened with the new API. --- src/analysis/special_functions/gaussian.lean | 17 +- src/analysis/special_functions/integrals.lean | 2 +- src/measure_theory/integral/exp_decay.lean | 28 +-- .../integral/integral_eq_improper.lean | 172 +++++++++++++++++- .../integral/interval_integral.lean | 16 +- 5 files changed, 191 insertions(+), 44 deletions(-) diff --git a/src/analysis/special_functions/gaussian.lean b/src/analysis/special_functions/gaussian.lean index edebd50124498..726f21901ca4a 100644 --- a/src/analysis/special_functions/gaussian.lean +++ b/src/analysis/special_functions/gaussian.lean @@ -158,25 +158,20 @@ lemma integral_mul_cexp_neg_mul_sq {b : ℂ} (hb : 0 < b.re) : ∫ r:ℝ in Ioi 0, (r : ℂ) * cexp (-b * r ^ 2) = (2 * b)⁻¹ := begin have hb' : b ≠ 0 := by { contrapose! hb, rw [hb, zero_re], }, - refine tendsto_nhds_unique (interval_integral_tendsto_integral_Ioi _ - (integrable_mul_cexp_neg_mul_sq hb).integrable_on filter.tendsto_id) _, have A : ∀ x:ℂ, has_deriv_at (λ x, - (2 * b)⁻¹ * cexp (-b * x^2)) (x * cexp (- b * x^2)) x, { intro x, convert (((has_deriv_at_pow 2 x)).const_mul (-b)).cexp.const_mul (- (2 * b)⁻¹) using 1, field_simp [hb'], ring }, - have : ∀ (y : ℝ), ∫ x in 0..(id y), ↑x * cexp (-b * x^2) - = (- (2 * b)⁻¹ * cexp (-b * y^2)) - (- (2 * b)⁻¹ * cexp (-b * 0^2)) := - λ y, interval_integral.integral_eq_sub_of_has_deriv_at - (λ x hx, (A x).comp_of_real) (integrable_mul_cexp_neg_mul_sq hb).interval_integrable, - simp_rw this, - have L : tendsto (λ (x : ℝ), (2 * b)⁻¹ - (2 * b)⁻¹ * cexp (-b * x ^ 2)) at_top - (𝓝 ((2 * b)⁻¹ - (2 * b)⁻¹ * 0)), - { refine tendsto_const_nhds.sub (tendsto.const_mul _ $ tendsto_zero_iff_norm_tendsto_zero.mpr _), + have B : tendsto (λ (y : ℝ), -(2 * b)⁻¹ * cexp (-b * ↑y ^ 2)) at_top (𝓝 (-(2 * b)⁻¹ * 0)), + { refine (tendsto.const_mul _ (tendsto_zero_iff_norm_tendsto_zero.mpr _)), simp_rw norm_cexp_neg_mul_sq b, exact tendsto_exp_at_bot.comp (tendsto.neg_const_mul_at_top (neg_lt_zero.2 hb) (tendsto_pow_at_top two_ne_zero)) }, - simpa using L, + convert integral_Ioi_of_has_deriv_at_of_tendsto' (λ x hx, (A ↑x).comp_of_real) + (integrable_mul_cexp_neg_mul_sq hb).integrable_on B, + simp only [mul_zero, of_real_zero, zero_pow', ne.def, bit0_eq_zero, nat.one_ne_zero, + not_false_iff, complex.exp_zero, mul_one, sub_neg_eq_add, zero_add], end /-- The *square* of the Gaussian integral `∫ x:ℝ, exp (-b * x^2)` is equal to `π / b`. -/ diff --git a/src/analysis/special_functions/integrals.lean b/src/analysis/special_functions/integrals.lean index 6bcc48d15b912..b7e245c98d162 100644 --- a/src/analysis/special_functions/integrals.lean +++ b/src/analysis/special_functions/integrals.lean @@ -67,7 +67,7 @@ begin have hderiv : ∀ x ∈ Ioo 0 c, has_deriv_at (λ x : ℝ, x ^ (r + 1) / (r + 1)) (x ^ r) x, { intros x hx, convert (real.has_deriv_at_rpow_const (or.inl hx.1.ne')).div_const (r + 1), field_simp [(by linarith : r + 1 ≠ 0)], ring, }, - apply integrable_on_deriv_of_nonneg hc _ hderiv, + apply integrable_on_deriv_of_nonneg _ hderiv, { intros x hx, apply rpow_nonneg_of_nonneg hx.1.le, }, { refine (continuous_on_id.rpow_const _).div_const _, intros x hx, right, linarith } }, intro c, rcases le_total 0 c with hc|hc, diff --git a/src/measure_theory/integral/exp_decay.lean b/src/measure_theory/integral/exp_decay.lean index 8f873fd96e193..2090e724c7a2d 100644 --- a/src/measure_theory/integral/exp_decay.lean +++ b/src/measure_theory/integral/exp_decay.lean @@ -18,33 +18,17 @@ for integrability: noncomputable theory open real interval_integral measure_theory set filter - -/-- Integral of `exp (-b * x)` over `(a, X)` is bounded as `X → ∞`. -/ -lemma integral_exp_neg_le {b : ℝ} (a X : ℝ) (h2 : 0 < b) : - (∫ x in a .. X, exp (-b * x)) ≤ exp (-b * a) / b := -begin - rw integral_deriv_eq_sub' (λ x, -exp (-b * x) / b), - -- goal 1/4: F(X) - F(a) is bounded - { simp only [tsub_le_iff_right], - rw [neg_div b (exp (-b * a)), neg_div b (exp (-b * X)), add_neg_self, neg_le, neg_zero], - exact (div_pos (exp_pos _) h2).le, }, - -- goal 2/4: the derivative of F is exp(-b x) - { ext1, simp [h2.ne'] }, - -- goal 3/4: F is differentiable - { intros x hx, simp [h2.ne'], }, - -- goal 4/4: exp(-b x) is continuous - { apply continuous.continuous_on, continuity } -end +open_locale topology /-- `exp (-b * x)` is integrable on `(a, ∞)`. -/ lemma exp_neg_integrable_on_Ioi (a : ℝ) {b : ℝ} (h : 0 < b) : integrable_on (λ x : ℝ, exp (-b * x)) (Ioi a) := begin - have : ∀ (X : ℝ), integrable_on (λ x : ℝ, exp (-b * x) ) (Ioc a X), - { intro X, exact (continuous_const.mul continuous_id).exp.integrable_on_Ioc }, - apply (integrable_on_Ioi_of_interval_integral_norm_bounded (exp (-b * a) / b) a this tendsto_id), - simp only [eventually_at_top, norm_of_nonneg (exp_pos _).le], - exact ⟨a, λ b2 hb2, integral_exp_neg_le a b2 h⟩, + have : tendsto (λ x, -exp (-b * x) / b) at_top (𝓝 (-0/b)), + { refine tendsto.div_const (tendsto.neg _) _, + exact tendsto_exp_at_bot.comp (tendsto_id.neg_const_mul_at_top (right.neg_neg_iff.2 h)) }, + refine integrable_on_Ioi_deriv_of_nonneg' (λ x hx, _) (λ x hx, (exp_pos _).le) this, + simpa [h.ne'] using ((has_deriv_at_id x).const_mul b).neg.exp.neg.div_const b, end /-- If `f` is continuous on `[a, ∞)`, and is `O (exp (-b * x))` at `∞` for some `b > 0`, then diff --git a/src/measure_theory/integral/integral_eq_improper.lean b/src/measure_theory/integral/integral_eq_improper.lean index 5f583e481dd94..335574962c38d 100644 --- a/src/measure_theory/integral/integral_eq_improper.lean +++ b/src/measure_theory/integral/integral_eq_improper.lean @@ -50,7 +50,15 @@ as an `ae_cover` w.r.t. `μ.restrict (Iic b)`, instead of using `(λ x, Ioc x b) then `∫ x in φ n, f x ∂μ` tends to `∫ x, f x ∂μ` as `n` tends to `+∞`. We then specialize these lemmas to various use cases involving intervals, which are frequent -in analysis. +in analysis. In particular, +- `measure_theory.integral_Ioi_of_has_deriv_at_of_tendsto` is a version of FTC-2 on the interval + `(a, +∞)`, giving the formula `∫ x in (a, +∞), g' x = l - g a` if `g'` is integrable and + `g` tends to `l` at `+∞`. +- `measure_theory.integral_Ioi_of_has_deriv_at_of_nonneg` gives the same result assuming that + `g'` is nonnegative instead of integrable. Its automatic integrability in this context is proved + in `measure_theory.integrable_on_Ioi_deriv_of_nonneg`. +- `measure_theory.integral_comp_smul_deriv_Ioi` is a version of the change of variables formula + on semi-infinite intervals. -/ open measure_theory filter set topological_space @@ -604,7 +612,7 @@ end lemma integrable_on_Ioc_of_interval_integral_norm_bounded_left {I a₀ b : ℝ} (hfi : ∀ i, integrable_on f $ Ioc (a i) b) (ha : tendsto a l $ 𝓝 a₀) - (h : ∀ᶠ i in l, (∫ x in Ioc (a i) b, ‖f x‖ ) ≤ I) : integrable_on f (Ioc a₀ b) := + (h : ∀ᶠ i in l, (∫ x in Ioc (a i) b, ‖f x‖) ≤ I) : integrable_on f (Ioc a₀ b) := integrable_on_Ioc_of_interval_integral_norm_bounded hfi ha tendsto_const_nhds h lemma integrable_on_Ioc_of_interval_integral_norm_bounded_right {I a b₀ : ℝ} @@ -660,12 +668,170 @@ end end integral_of_interval_integral +open real +open_locale interval + +section Ioi_FTC + +variables {E : Type*} {f f' : ℝ → E} {g g' : ℝ → ℝ} {a b l : ℝ} {m : E} + [normed_add_comm_group E] [normed_space ℝ E] [complete_space E] + +/-- **Fundamental theorem of calculus-2**, on semi-infinite intervals `(a, +∞)`. +When a function has a limit at infinity `m`, and its derivative is integrable, then the +integral of the derivative on `(a, +∞)` is `m - f a`. Version assuming differentiability +on `(a, +∞)` and continuity on `[a, +∞)`.-/ +lemma integral_Ioi_of_has_deriv_at_of_tendsto (hcont : continuous_on f (Ici a)) + (hderiv : ∀ x ∈ Ioi a, has_deriv_at f (f' x) x) + (f'int : integrable_on f' (Ioi a)) (hf : tendsto f at_top (𝓝 m)) : + ∫ x in Ioi a, f' x = m - f a := +begin + refine tendsto_nhds_unique (interval_integral_tendsto_integral_Ioi a f'int tendsto_id) _, + apply tendsto.congr' _ (hf.sub_const _), + filter_upwards [Ioi_mem_at_top a] with x hx, + have h'x : a ≤ id x := le_of_lt hx, + symmetry, + apply interval_integral.integral_eq_sub_of_has_deriv_at_of_le h'x + (hcont.mono Icc_subset_Ici_self) (λ y hy, hderiv y hy.1), + rw interval_integrable_iff_integrable_Ioc_of_le h'x, + exact f'int.mono (λ y hy, hy.1) le_rfl, +end + +/-- **Fundamental theorem of calculus-2**, on semi-infinite intervals `(a, +∞)`. +When a function has a limit at infinity `m`, and its derivative is integrable, then the +integral of the derivative on `(a, +∞)` is `m - f a`. Version assuming differentiability +on `[a, +∞)`. -/ +lemma integral_Ioi_of_has_deriv_at_of_tendsto' + (hderiv : ∀ x ∈ Ici a, has_deriv_at f (f' x) x) + (f'int : integrable_on f' (Ioi a)) (hf : tendsto f at_top (𝓝 m)) : + ∫ x in Ioi a, f' x = m - f a := +begin + apply integral_Ioi_of_has_deriv_at_of_tendsto _ (λ x hx, hderiv x (le_of_lt hx)) f'int hf, + assume x hx, + exact (hderiv x hx).continuous_at.continuous_within_at, +end + +/-- When a function has a limit at infinity, and its derivative is nonnegative, then the derivative +is automatically integrable on `(a, +∞)`. Version assuming differentiability +on `(a, +∞)` and continuity on `[a, +∞)`. -/ +lemma integrable_on_Ioi_deriv_of_nonneg (hcont : continuous_on g (Ici a)) + (hderiv : ∀ x ∈ Ioi a, has_deriv_at g (g' x) x) + (g'pos : ∀ x ∈ Ioi a, 0 ≤ g' x) (hg : tendsto g at_top (𝓝 l)) : + integrable_on g' (Ioi a) := +begin + apply integrable_on_Ioi_of_interval_integral_norm_tendsto (l - g a) a (λ x, _) tendsto_id, swap, + { exact interval_integral.integrable_on_deriv_of_nonneg (hcont.mono Icc_subset_Ici_self) + (λ y hy, hderiv y hy.1) (λ y hy, g'pos y hy.1) }, + apply tendsto.congr' _ (hg.sub_const _), + filter_upwards [Ioi_mem_at_top a] with x hx, + have h'x : a ≤ id x := le_of_lt hx, + calc g x - g a = ∫ y in a..id x, g' y : + begin + symmetry, + apply interval_integral.integral_eq_sub_of_has_deriv_at_of_le h'x + (hcont.mono Icc_subset_Ici_self) (λ y hy, hderiv y hy.1), + rw interval_integrable_iff_integrable_Ioc_of_le h'x, + exact interval_integral.integrable_on_deriv_of_nonneg (hcont.mono Icc_subset_Ici_self) + (λ y hy, hderiv y hy.1) (λ y hy, g'pos y hy.1) + end + ... = ∫ y in a..id x, ‖g' y‖ : + begin + simp_rw interval_integral.integral_of_le h'x, + refine set_integral_congr (measurable_set_Ioc) (λ y hy, _), + dsimp, + rw abs_of_nonneg, + exact g'pos _ hy.1, + end +end + +/-- When a function has a limit at infinity, and its derivative is nonnegative, then the derivative +is automatically integrable on `(a, +∞)`. Version assuming differentiability +on `[a, +∞)`. -/ +lemma integrable_on_Ioi_deriv_of_nonneg' + (hderiv : ∀ x ∈ Ici a, has_deriv_at g (g' x) x) + (g'pos : ∀ x ∈ Ioi a, 0 ≤ g' x) (hg : tendsto g at_top (𝓝 l)) : + integrable_on g' (Ioi a) := +begin + apply integrable_on_Ioi_deriv_of_nonneg _ (λ x hx, hderiv x (le_of_lt hx)) g'pos hg, + assume x hx, + exact (hderiv x hx).continuous_at.continuous_within_at, +end + +/-- When a function has a limit at infinity `l`, and its derivative is nonnegative, then the +integral of the derivative on `(a, +∞)` is `l - g a` (and the derivative is integrable, see +`integrable_on_Ioi_deriv_of_nonneg`). Version assuming differentiability on `(a, +∞)` and +continuity on `[a, +∞)`. -/ +lemma integral_Ioi_of_has_deriv_at_of_nonneg (hcont : continuous_on g (Ici a)) + (hderiv : ∀ x ∈ Ioi a, has_deriv_at g (g' x) x) + (g'pos : ∀ x ∈ Ioi a, 0 ≤ g' x) (hg : tendsto g at_top (𝓝 l)) : + ∫ x in Ioi a, g' x = l - g a := +integral_Ioi_of_has_deriv_at_of_tendsto hcont hderiv + (integrable_on_Ioi_deriv_of_nonneg hcont hderiv g'pos hg) hg + +/-- When a function has a limit at infinity `l`, and its derivative is nonnegative, then the +integral of the derivative on `(a, +∞)` is `l - g a` (and the derivative is integrable, see +`integrable_on_Ioi_deriv_of_nonneg'`). Version assuming differentiability on `[a, +∞)`. -/ +lemma integral_Ioi_of_has_deriv_at_of_nonneg' + (hderiv : ∀ x ∈ Ici a, has_deriv_at g (g' x) x) + (g'pos : ∀ x ∈ Ioi a, 0 ≤ g' x) (hg : tendsto g at_top (𝓝 l)) : + ∫ x in Ioi a, g' x = l - g a := +integral_Ioi_of_has_deriv_at_of_tendsto' hderiv + (integrable_on_Ioi_deriv_of_nonneg' hderiv g'pos hg) hg + +/-- When a function has a limit at infinity, and its derivative is nonpositive, then the derivative +is automatically integrable on `(a, +∞)`. Version assuming differentiability +on `(a, +∞)` and continuity on `[a, +∞)`. -/ +lemma integrable_on_Ioi_deriv_of_nonpos (hcont : continuous_on g (Ici a)) + (hderiv : ∀ x ∈ Ioi a, has_deriv_at g (g' x) x) + (g'neg : ∀ x ∈ Ioi a, g' x ≤ 0) (hg : tendsto g at_top (𝓝 l)) : + integrable_on g' (Ioi a) := +begin + apply integrable_neg_iff.1, + exact integrable_on_Ioi_deriv_of_nonneg hcont.neg (λ x hx, (hderiv x hx).neg) + (λ x hx, neg_nonneg_of_nonpos (g'neg x hx)) hg.neg, +end + +/-- When a function has a limit at infinity, and its derivative is nonpositive, then the derivative +is automatically integrable on `(a, +∞)`. Version assuming differentiability +on `[a, +∞)`. -/ +lemma integrable_on_Ioi_deriv_of_nonpos' + (hderiv : ∀ x ∈ Ici a, has_deriv_at g (g' x) x) + (g'neg : ∀ x ∈ Ioi a, g' x ≤ 0) (hg : tendsto g at_top (𝓝 l)) : + integrable_on g' (Ioi a) := +begin + apply integrable_on_Ioi_deriv_of_nonpos _ (λ x hx, hderiv x (le_of_lt hx)) g'neg hg, + assume x hx, + exact (hderiv x hx).continuous_at.continuous_within_at, +end + +/-- When a function has a limit at infinity `l`, and its derivative is nonpositive, then the +integral of the derivative on `(a, +∞)` is `l - g a` (and the derivative is integrable, see +`integrable_on_Ioi_deriv_of_nonneg`). Version assuming differentiability on `(a, +∞)` and +continuity on `[a, +∞)`. -/ +lemma integral_Ioi_of_has_deriv_at_of_nonpos (hcont : continuous_on g (Ici a)) + (hderiv : ∀ x ∈ Ioi a, has_deriv_at g (g' x) x) + (g'neg : ∀ x ∈ Ioi a, g' x ≤ 0) (hg : tendsto g at_top (𝓝 l)) : + ∫ x in Ioi a, g' x = l - g a := +integral_Ioi_of_has_deriv_at_of_tendsto hcont hderiv + (integrable_on_Ioi_deriv_of_nonpos hcont hderiv g'neg hg) hg + +/-- When a function has a limit at infinity `l`, and its derivative is nonpositive, then the +integral of the derivative on `(a, +∞)` is `l - g a` (and the derivative is integrable, see +`integrable_on_Ioi_deriv_of_nonneg'`). Version assuming differentiability on `[a, +∞)`. -/ +lemma integral_Ioi_of_has_deriv_at_of_nonpos' + (hderiv : ∀ x ∈ Ici a, has_deriv_at g (g' x) x) + (g'neg : ∀ x ∈ Ioi a, g' x ≤ 0) (hg : tendsto g at_top (𝓝 l)) : + ∫ x in Ioi a, g' x = l - g a := +integral_Ioi_of_has_deriv_at_of_tendsto' hderiv + (integrable_on_Ioi_deriv_of_nonpos' hderiv g'neg hg) hg + +end Ioi_FTC + section Ioi_change_variables open real open_locale interval -variables {E : Type*} {μ : measure ℝ} {f : ℝ → E} +variables {E : Type*} {f : ℝ → E} [normed_add_comm_group E] [normed_space ℝ E] [complete_space E] /-- Change-of-variables formula for `Ioi` integrals of vector-valued functions, proved by taking diff --git a/src/measure_theory/integral/interval_integral.lean b/src/measure_theory/integral/interval_integral.lean index 6b96f2e9d66d4..a9addd4c9b351 100644 --- a/src/measure_theory/integral/interval_integral.lean +++ b/src/measure_theory/integral/interval_integral.lean @@ -2524,11 +2524,13 @@ end -/ /-- When the right derivative of a function is nonnegative, then it is automatically integrable. -/ -lemma integrable_on_deriv_right_of_nonneg (hab : a ≤ b) (hcont : continuous_on g (Icc a b)) +lemma integrable_on_deriv_right_of_nonneg (hcont : continuous_on g (Icc a b)) (hderiv : ∀ x ∈ Ioo a b, has_deriv_within_at g (g' x) (Ioi x) x) (g'pos : ∀ x ∈ Ioo a b, 0 ≤ g' x) : integrable_on g' (Ioc a b) := begin + by_cases hab : a < b, swap, + { simp [Ioc_eq_empty hab] }, rw integrable_on_Ioc_iff_integrable_on_Ioo, have meas_g' : ae_measurable g' (volume.restrict (Ioo a b)), { apply (ae_measurable_deriv_within_Ioi g _).congr, @@ -2549,8 +2551,8 @@ begin lintegral_coe_eq_integral _ intF, rw A at hf, have B : ∫ (x : ℝ) in Ioo a b, F x ≤ g b - g a, - { rw [← integral_Ioc_eq_integral_Ioo, ← interval_integral.integral_of_le hab], - apply integral_le_sub_of_has_deriv_right_of_le hab hcont hderiv _ (λ x hx, _), + { rw [← integral_Ioc_eq_integral_Ioo, ← interval_integral.integral_of_le hab.le], + apply integral_le_sub_of_has_deriv_right_of_le hab.le hcont hderiv _ (λ x hx, _), { rwa integrable_on_Icc_iff_integrable_on_Ioo }, { convert nnreal.coe_le_coe.2 (fle x), simp only [real.norm_of_nonneg (g'pos x hx), coe_nnnorm] } }, @@ -2559,11 +2561,11 @@ end /-- When the derivative of a function is nonnegative, then it is automatically integrable, Ioc version. -/ -lemma integrable_on_deriv_of_nonneg (hab : a ≤ b) (hcont : continuous_on g (Icc a b)) +lemma integrable_on_deriv_of_nonneg (hcont : continuous_on g (Icc a b)) (hderiv : ∀ x ∈ Ioo a b, has_deriv_at g (g' x) x) (g'pos : ∀ x ∈ Ioo a b, 0 ≤ g' x) : integrable_on g' (Ioc a b) := -integrable_on_deriv_right_of_nonneg hab hcont (λ x hx, (hderiv x hx).has_deriv_within_at) g'pos +integrable_on_deriv_right_of_nonneg hcont (λ x hx, (hderiv x hx).has_deriv_within_at) g'pos /-- When the derivative of a function is nonnegative, then it is automatically integrable, interval version. -/ @@ -2575,10 +2577,10 @@ begin cases le_total a b with hab hab, { simp only [uIcc_of_le, min_eq_left, max_eq_right, hab, interval_integrable, hab, Ioc_eq_empty_of_le, integrable_on_empty, and_true] at hcont hderiv hpos ⊢, - exact integrable_on_deriv_of_nonneg hab hcont hderiv hpos, }, + exact integrable_on_deriv_of_nonneg hcont hderiv hpos, }, { simp only [uIcc_of_ge, min_eq_right, max_eq_left, hab, interval_integrable, Ioc_eq_empty_of_le, integrable_on_empty, true_and] at hcont hderiv hpos ⊢, - exact integrable_on_deriv_of_nonneg hab hcont hderiv hpos } + exact integrable_on_deriv_of_nonneg hcont hderiv hpos } end /-! From 7dfe85833014fb54258a228081ebb76b7e96ec98 Mon Sep 17 00:00:00 2001 From: Floris van Doorn Date: Sat, 22 Apr 2023 07:09:42 +0000 Subject: [PATCH 0861/1291] feat(geometry/manifold/vector_bundle/tangent): some equations about trivializations (#18825) * From the sphere eversion project --- .../manifold/vector_bundle/tangent.lean | 43 ++++++++++++++++++- src/topology/vector_bundle/basic.lean | 38 ++++++++++++++-- 2 files changed, 76 insertions(+), 5 deletions(-) diff --git a/src/geometry/manifold/vector_bundle/tangent.lean b/src/geometry/manifold/vector_bundle/tangent.lean index 870a40386d099..a707c965cf24e 100644 --- a/src/geometry/manifold/vector_bundle/tangent.lean +++ b/src/geometry/manifold/vector_bundle/tangent.lean @@ -141,11 +141,10 @@ does not pick wrong instances. In this section, we record the right instances fo them, noting in particular that the tangent bundle is a smooth manifold. -/ section -local attribute [reducible] tangent_space variables {M} (x : M) -instance : module 𝕜 (tangent_space I x) := by apply_instance +instance : module 𝕜 (tangent_space I x) := by delta_instance tangent_space instance : inhabited (tangent_space I x) := ⟨0⟩ end @@ -219,6 +218,41 @@ by simp only [fiber_bundle.charted_space_chart_at, and_iff_left_iff_imp] with mf @[simp, mfld_simps] lemma coe_chart_at_symm_fst (p : H × E) (q : TM) : ((chart_at (model_prod H E) q).symm p).1 = ((chart_at H q.1).symm : H → M) p.1 := rfl +@[simp, mfld_simps] lemma trivialization_at_continuous_linear_map_at {b₀ b : M} + (hb : b ∈ (trivialization_at E (tangent_space I) b₀).base_set) : + (trivialization_at E (tangent_space I) b₀).continuous_linear_map_at 𝕜 b = + (tangent_bundle_core I M).coord_change (achart H b) (achart H b₀) b := +(tangent_bundle_core I M).local_triv_continuous_linear_map_at hb + +@[simp, mfld_simps] lemma trivialization_at_symmL {b₀ b : M} + (hb : b ∈ (trivialization_at E (tangent_space I) b₀).base_set) : + (trivialization_at E (tangent_space I) b₀).symmL 𝕜 b = + (tangent_bundle_core I M).coord_change (achart H b₀) (achart H b) b := +(tangent_bundle_core I M).local_triv_symmL hb + +@[simp, mfld_simps] +lemma coord_change_model_space (b b' x : F) : + (tangent_bundle_core 𝓘(𝕜, F) F).coord_change (achart F b) (achart F b') x = 1 := +by simpa only [tangent_bundle_core_coord_change] with mfld_simps using + fderiv_within_id unique_diff_within_at_univ + +@[simp, mfld_simps] +lemma symmL_model_space (b b' : F) : + (trivialization_at F (tangent_space 𝓘(𝕜, F)) b).symmL 𝕜 b' = (1 : F →L[𝕜] F) := +begin + rw [tangent_bundle.trivialization_at_symmL, coord_change_model_space], + apply mem_univ +end + +@[simp, mfld_simps] +lemma continuous_linear_map_at_model_space (b b' : F) : + (trivialization_at F (tangent_space 𝓘(𝕜, F)) b).continuous_linear_map_at 𝕜 b' = + (1 : F →L[𝕜] F) := +begin + rw [tangent_bundle.trivialization_at_continuous_linear_map_at, coord_change_model_space], + apply mem_univ +end + end tangent_bundle instance tangent_bundle_core.is_smooth : (tangent_bundle_core I M).is_smooth I := @@ -270,6 +304,11 @@ by { unfold_coes, simp_rw [tangent_bundle_model_space_chart_at], refl } by { unfold_coes, simp_rw [local_homeomorph.symm_to_local_equiv, tangent_bundle_model_space_chart_at], refl } +lemma tangent_bundle_core_coord_change_model_space (x x' z : H) : + (tangent_bundle_core I H).coord_change (achart H x) (achart H x') z = + continuous_linear_map.id 𝕜 E := +by { ext v, exact (tangent_bundle_core I H).coord_change_self (achart _ z) z (mem_univ _) v } + variable (H) /-- The canonical identification between the tangent bundle to the model space and the product, as a homeomorphism -/ diff --git a/src/topology/vector_bundle/basic.lean b/src/topology/vector_bundle/basic.lean index 235477250dbab..364f13df75491 100644 --- a/src/topology/vector_bundle/basic.lean +++ b/src/topology/vector_bundle/basic.lean @@ -43,7 +43,7 @@ noncomputable theory open bundle set open_locale classical bundle -variables (R 𝕜 : Type*) {B : Type*} (F : Type*) (E : B → Type*) +variables (R : Type*) {B : Type*} (F : Type*) (E : B → Type*) section topological_vector_space variables {B F E} [semiring R] @@ -560,13 +560,13 @@ instance add_comm_group_fiber [add_comm_group F] : ∀ (x : B), add_comm_group ( by dsimp [vector_bundle_core.fiber]; delta_instance fiber_bundle_core.fiber /-- The projection from the total space of a fiber bundle core, on its base. -/ -@[reducible, simp, mfld_simps] def proj : total_space Z.fiber → B := total_space.proj +@[reducible, simp, mfld_simps] protected def proj : total_space Z.fiber → B := total_space.proj /-- The total space of the vector bundle, as a convenience function for dot notation. It is by definition equal to `bundle.total_space Z.fiber`, a.k.a. `Σ x, Z.fiber x` but with a different name for typeclass inference. -/ @[nolint unused_arguments, reducible] -def total_space := bundle.total_space Z.fiber +protected def total_space := bundle.total_space Z.fiber /-- Local homeomorphism version of the trivialization change. -/ def triv_change (i j : ι) : local_homeomorph (B × F) (B × F) := @@ -675,6 +675,38 @@ fiber_bundle_core.continuous_proj Z lemma is_open_map_proj : is_open_map Z.proj := fiber_bundle_core.is_open_map_proj Z +variables {i j} + +@[simp, mfld_simps] lemma local_triv_continuous_linear_map_at {b : B} (hb : b ∈ Z.base_set i) : + (Z.local_triv i).continuous_linear_map_at R b = Z.coord_change (Z.index_at b) i b := +begin + ext1 v, + rw [(Z.local_triv i).continuous_linear_map_at_apply R, (Z.local_triv i).coe_linear_map_at_of_mem], + exacts [rfl, hb] +end + +@[simp, mfld_simps] lemma trivialization_at_continuous_linear_map_at {b₀ b : B} + (hb : b ∈ (trivialization_at F Z.fiber b₀).base_set) : + (trivialization_at F Z.fiber b₀).continuous_linear_map_at R b = + Z.coord_change (Z.index_at b) (Z.index_at b₀) b := +Z.local_triv_continuous_linear_map_at hb + +@[simp, mfld_simps] lemma local_triv_symmL {b : B} (hb : b ∈ Z.base_set i) : + (Z.local_triv i).symmL R b = Z.coord_change i (Z.index_at b) b := +by { ext1 v, rw [(Z.local_triv i).symmL_apply R, (Z.local_triv i).symm_apply], exacts [rfl, hb] } + +@[simp, mfld_simps] lemma trivialization_at_symmL {b₀ b : B} + (hb : b ∈ (trivialization_at F Z.fiber b₀).base_set) : + (trivialization_at F Z.fiber b₀).symmL R b = Z.coord_change (Z.index_at b₀) (Z.index_at b) b := +Z.local_triv_symmL hb + +@[simp, mfld_simps] lemma trivialization_at_coord_change_eq {b₀ b₁ b : B} + (hb : b ∈ (trivialization_at F Z.fiber b₀).base_set ∩ (trivialization_at F Z.fiber b₁).base_set) + (v : F) : + (trivialization_at F Z.fiber b₀).coord_changeL R (trivialization_at F Z.fiber b₁) b v = + Z.coord_change (Z.index_at b₀) (Z.index_at b₁) b v := +Z.local_triv_coord_change_eq _ _ hb v + end vector_bundle_core end From 4f81bc21e32048db7344b7867946e992cf5f68cc Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Sat, 22 Apr 2023 08:22:26 +0000 Subject: [PATCH 0862/1291] chore(*): add mathlib4 synchronization comments (#18851) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.monoid_algebra.ideal` * `algebra.ring_quot` * `algebraic_topology.dold_kan.decomposition` * `algebraic_topology.dold_kan.degeneracies` * `algebraic_topology.dold_kan.n_reflects_iso` * `algebraic_topology.dold_kan.split_simplicial_object` * `category_theory.adjunction.over` * `data.polynomial.module` * `linear_algebra.affine_space.independent` * `linear_algebra.matrix.adjugate` * `linear_algebra.matrix.nondegenerate` * `ring_theory.mv_polynomial.ideal` --- src/algebra/monoid_algebra/ideal.lean | 3 +++ src/algebra/ring_quot.lean | 3 +++ src/algebraic_topology/dold_kan/decomposition.lean | 3 +++ src/algebraic_topology/dold_kan/degeneracies.lean | 3 +++ src/algebraic_topology/dold_kan/n_reflects_iso.lean | 3 +++ src/algebraic_topology/dold_kan/split_simplicial_object.lean | 3 +++ src/category_theory/adjunction/over.lean | 3 +++ src/data/polynomial/module.lean | 3 +++ src/linear_algebra/affine_space/independent.lean | 3 +++ src/linear_algebra/matrix/adjugate.lean | 3 +++ src/linear_algebra/matrix/nondegenerate.lean | 3 +++ src/ring_theory/mv_polynomial/ideal.lean | 3 +++ 12 files changed, 36 insertions(+) diff --git a/src/algebra/monoid_algebra/ideal.lean b/src/algebra/monoid_algebra/ideal.lean index 046f3dff58f7f..c5b235ec0c0db 100644 --- a/src/algebra/monoid_algebra/ideal.lean +++ b/src/algebra/monoid_algebra/ideal.lean @@ -9,6 +9,9 @@ import ring_theory.ideal.basic /-! # Lemmas about ideals of `monoid_algebra` and `add_monoid_algebra` + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ variables {k A G : Type*} diff --git a/src/algebra/ring_quot.lean b/src/algebra/ring_quot.lean index e153968e3fcfb..026ee392e01c4 100644 --- a/src/algebra/ring_quot.lean +++ b/src/algebra/ring_quot.lean @@ -9,6 +9,9 @@ import ring_theory.ideal.quotient /-! # Quotients of non-commutative rings +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Unfortunately, ideals have only been developed in the commutative case as `ideal`, and it's not immediately clear how one should formalise ideals in the non-commutative case. diff --git a/src/algebraic_topology/dold_kan/decomposition.lean b/src/algebraic_topology/dold_kan/decomposition.lean index fa1471a071589..1d89283129af6 100644 --- a/src/algebraic_topology/dold_kan/decomposition.lean +++ b/src/algebraic_topology/dold_kan/decomposition.lean @@ -10,6 +10,9 @@ import algebraic_topology.dold_kan.p_infty # Decomposition of the Q endomorphisms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we obtain a lemma `decomposition_Q` which expresses explicitly the projection `(Q q).f (n+1) : X _[n+1] ⟶ X _[n+1]` (`X : simplicial_object C` with `C` a preadditive category) as diff --git a/src/algebraic_topology/dold_kan/degeneracies.lean b/src/algebraic_topology/dold_kan/degeneracies.lean index 35aac95468ce1..c3b3de9892ba2 100644 --- a/src/algebraic_topology/dold_kan/degeneracies.lean +++ b/src/algebraic_topology/dold_kan/degeneracies.lean @@ -11,6 +11,9 @@ import tactic.fin_cases # Behaviour of P_infty with respect to degeneracies +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + For any `X : simplicial_object C` where `C` is an abelian category, the projector `P_infty : K[X] ⟶ K[X]` is supposed to be the projection on the normalized subcomplex, parallel to the degenerate subcomplex, i.e. diff --git a/src/algebraic_topology/dold_kan/n_reflects_iso.lean b/src/algebraic_topology/dold_kan/n_reflects_iso.lean index 1dedf74b1765c..d9cd2e612a1e0 100644 --- a/src/algebraic_topology/dold_kan/n_reflects_iso.lean +++ b/src/algebraic_topology/dold_kan/n_reflects_iso.lean @@ -13,6 +13,9 @@ import category_theory.idempotents.karoubi_karoubi # N₁ and N₂ reflects isomorphisms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, it is shown that the functors `N₁ : simplicial_object C ⥤ karoubi (chain_complex C ℕ)` and `N₂ : karoubi (simplicial_object C) ⥤ karoubi (chain_complex C ℕ))` diff --git a/src/algebraic_topology/dold_kan/split_simplicial_object.lean b/src/algebraic_topology/dold_kan/split_simplicial_object.lean index bbb6ac4e24967..f3b2c7fec482f 100644 --- a/src/algebraic_topology/dold_kan/split_simplicial_object.lean +++ b/src/algebraic_topology/dold_kan/split_simplicial_object.lean @@ -12,6 +12,9 @@ import algebraic_topology.dold_kan.functor_n # Split simplicial objects in preadditive categories +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define a functor `nondeg_complex : simplicial_object.split C ⥤ chain_complex C ℕ` when `C` is a preadditive category with finite coproducts, and get an isomorphism `to_karoubi_nondeg_complex_iso_N₁ : nondeg_complex ⋙ to_karoubi _ ≅ forget C ⋙ dold_kan.N₁`. diff --git a/src/category_theory/adjunction/over.lean b/src/category_theory/adjunction/over.lean index 048eb4e5ac321..d709fed80a217 100644 --- a/src/category_theory/adjunction/over.lean +++ b/src/category_theory/adjunction/over.lean @@ -10,6 +10,9 @@ import category_theory.over /-! # Adjunctions related to the over category +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Construct the left adjoint `star X` to `over.forget X : over X ⥤ C`. ## TODO diff --git a/src/data/polynomial/module.lean b/src/data/polynomial/module.lean index f0e402e31c0b7..d9e8cc96046d1 100644 --- a/src/data/polynomial/module.lean +++ b/src/data/polynomial/module.lean @@ -9,6 +9,9 @@ import ring_theory.finite_type /-! # Polynomial module +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we define the polynomial module for an `R`-module `M`, i.e. the `R[X]`-module `M[X]`. This is defined as an type alias `polynomial_module R M := ℕ →₀ M`, since there might be different diff --git a/src/linear_algebra/affine_space/independent.lean b/src/linear_algebra/affine_space/independent.lean index e781e031dfef7..7918c3ff4d4a8 100644 --- a/src/linear_algebra/affine_space/independent.lean +++ b/src/linear_algebra/affine_space/independent.lean @@ -13,6 +13,9 @@ import linear_algebra.basis /-! # Affine independence +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines affinely independent families of points. ## Main definitions diff --git a/src/linear_algebra/matrix/adjugate.lean b/src/linear_algebra/matrix/adjugate.lean index a3868b6d8d3f7..eb87b9f7f57c4 100644 --- a/src/linear_algebra/matrix/adjugate.lean +++ b/src/linear_algebra/matrix/adjugate.lean @@ -11,6 +11,9 @@ import ring_theory.polynomial.basic /-! # Cramer's rule and adjugate matrices +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The adjugate matrix is the transpose of the cofactor matrix. It is calculated with Cramer's rule, which we introduce first. The vectors returned by Cramer's rule are given by the linear map `cramer`, diff --git a/src/linear_algebra/matrix/nondegenerate.lean b/src/linear_algebra/matrix/nondegenerate.lean index 5dfbf769f0461..9c76a186f0680 100644 --- a/src/linear_algebra/matrix/nondegenerate.lean +++ b/src/linear_algebra/matrix/nondegenerate.lean @@ -10,6 +10,9 @@ import linear_algebra.matrix.adjugate /-! # Matrices associated with non-degenerate bilinear forms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions * `matrix.nondegenerate A`: the proposition that when interpreted as a bilinear form, the matrix `A` diff --git a/src/ring_theory/mv_polynomial/ideal.lean b/src/ring_theory/mv_polynomial/ideal.lean index 2faf6d7280dcf..b7ddbc120d4e0 100644 --- a/src/ring_theory/mv_polynomial/ideal.lean +++ b/src/ring_theory/mv_polynomial/ideal.lean @@ -9,6 +9,9 @@ import data.mv_polynomial.division /-! # Lemmas about ideals of `mv_polynomial` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Notably this contains results about monomial ideals. ## Main results From 323b7f2616426313505aae4e09ffbea6013862fc Mon Sep 17 00:00:00 2001 From: Moritz Doll Date: Sat, 22 Apr 2023 13:50:32 +0000 Subject: [PATCH 0863/1291] feat(analysis/calculus/cont_diff): iterated derivatives of maps into clm's (#18756) --- src/analysis/calculus/cont_diff.lean | 71 ++++++++++++++++++++++++++++ 1 file changed, 71 insertions(+) diff --git a/src/analysis/calculus/cont_diff.lean b/src/analysis/calculus/cont_diff.lean index 5ef0d2d58d2a5..790054897f53b 100644 --- a/src/analysis/calculus/cont_diff.lean +++ b/src/analysis/calculus/cont_diff.lean @@ -166,6 +166,21 @@ by { rw [subsingleton.elim f (λ _, 0)], exact cont_diff_within_at_const } cont_diff_on 𝕜 n f s := by { rw [subsingleton.elim f (λ _, 0)], exact cont_diff_on_const } +lemma iterated_fderiv_succ_const (n : ℕ) (c : F) : iterated_fderiv 𝕜 (n + 1) (λ (y : E), c) = 0 := +begin + ext x m, + simp only [iterated_fderiv_succ_apply_right, fderiv_const, pi.zero_apply, + iterated_fderiv_zero_fun, continuous_multilinear_map.zero_apply, + continuous_linear_map.zero_apply], +end + +lemma iterated_fderiv_const_of_ne {n : ℕ} (hn : n ≠ 0) (c : F) : + iterated_fderiv 𝕜 n (λ (y : E), c) = 0 := +begin + cases nat.exists_eq_succ_of_ne_zero hn with k hk, + rw [hk, iterated_fderiv_succ_const], +end + /-! ### Smoothness of linear functions -/ /-- @@ -2602,3 +2617,59 @@ begin end end + +section apply + +lemma norm_iterated_fderiv_within_clm_apply {f : E → (F →L[𝕜] G)} {g : E → F} {s : set E} {x : E} + {N : ℕ∞} {n : ℕ} (hf : cont_diff_on 𝕜 N f s) (hg : cont_diff_on 𝕜 N g s) (hs : unique_diff_on 𝕜 s) + (hx : x ∈ s) (hn : ↑n ≤ N) : + ‖iterated_fderiv_within 𝕜 n (λ y, (f y) (g y)) s x‖ ≤ + (finset.range (n + 1)).sum (λ i, ↑(n.choose i) * ‖iterated_fderiv_within 𝕜 i f s x‖ * + ‖iterated_fderiv_within 𝕜 (n - i) g s x‖) := +begin + let B : (F →L[𝕜] G) →L[𝕜] F →L[𝕜] G := + continuous_linear_map.flip (continuous_linear_map.apply 𝕜 G), + have hB : ‖B‖ ≤ 1 := + begin + simp only [continuous_linear_map.op_norm_flip, continuous_linear_map.apply], + refine continuous_linear_map.op_norm_le_bound _ zero_le_one (λ f, _), + simp only [continuous_linear_map.coe_id', id.def, one_mul], + end, + exact B.norm_iterated_fderiv_within_le_of_bilinear_of_le_one hf hg hs hx hn hB, +end + +lemma norm_iterated_fderiv_clm_apply {f : E → (F →L[𝕜] G)} {g : E → F} + {N : ℕ∞} {n : ℕ} (hf : cont_diff 𝕜 N f) (hg : cont_diff 𝕜 N g) (x : E) (hn : ↑n ≤ N): + ‖iterated_fderiv 𝕜 n (λ (y : E), (f y) (g y)) x‖ ≤ + (finset.range (n + 1)).sum (λ (i : ℕ), ↑(n.choose i) * ‖iterated_fderiv 𝕜 i f x‖ * + ‖iterated_fderiv 𝕜 (n - i) g x‖) := +begin + simp only [← iterated_fderiv_within_univ], + exact norm_iterated_fderiv_within_clm_apply hf.cont_diff_on hg.cont_diff_on unique_diff_on_univ + (set.mem_univ x) hn, +end + +lemma norm_iterated_fderiv_within_clm_apply_const {f : E → (F →L[𝕜] G)} {c : F} {s : set E} {x : E} + {N : ℕ∞} {n : ℕ} (hf : cont_diff_on 𝕜 N f s) (hs : unique_diff_on 𝕜 s) (hx : x ∈ s) + (hn : ↑n ≤ N) : ‖iterated_fderiv_within 𝕜 n (λ (y : E), (f y) c) s x‖ ≤ + ‖c‖ * ‖iterated_fderiv_within 𝕜 n f s x‖ := +begin + let g : (F →L[𝕜] G) →L[𝕜] G := continuous_linear_map.apply 𝕜 G c, + have h := g.norm_comp_continuous_multilinear_map_le (iterated_fderiv_within 𝕜 n f s x), + rw ← g.iterated_fderiv_within_comp_left hf hs hx hn at h, + refine h.trans (mul_le_mul_of_nonneg_right _ (norm_nonneg _)), + refine g.op_norm_le_bound (norm_nonneg _) (λ f, _), + rw [continuous_linear_map.apply_apply, mul_comm], + exact f.le_op_norm c, +end + +lemma norm_iterated_fderiv_clm_apply_const {f : E → (F →L[𝕜] G)} {c : F} {x : E} {N : ℕ∞} {n : ℕ} + (hf : cont_diff 𝕜 N f) (hn : ↑n ≤ N) : + ‖iterated_fderiv 𝕜 n (λ (y : E), (f y) c) x‖ ≤ ‖c‖ * ‖iterated_fderiv 𝕜 n f x‖ := +begin + simp only [← iterated_fderiv_within_univ], + refine norm_iterated_fderiv_within_clm_apply_const hf.cont_diff_on unique_diff_on_univ + (set.mem_univ x) hn, +end + +end apply From 7e281deff072232a3c5b3e90034bd65dde396312 Mon Sep 17 00:00:00 2001 From: damiano Date: Sat, 22 Apr 2023 15:18:24 +0000 Subject: [PATCH 0864/1291] feat(topology/extremally_disconnected): Extremally disconnected topological spaces (#8196) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit From LTE Author: Johan Commelin [Zulip discussion](https://leanprover.zulipchat.com/#narrow/stream/267928-condensed-mathematics/topic/extremally.20disconnected.20sets) Co-authored-by: Patrick Massot Co-authored-by: Eric Wieser Co-authored-by: Yaël Dillies --- docs/references.bib | 15 +++ src/topology/extremally_disconnected.lean | 114 ++++++++++++++++++++++ 2 files changed, 129 insertions(+) create mode 100644 src/topology/extremally_disconnected.lean diff --git a/docs/references.bib b/docs/references.bib index cd493023456fa..da809a2690cf7 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -999,6 +999,21 @@ @Book{ GierzEtAl1980 mrreviewer = {James W. Lea, Jr.} } +@Article{ gleason1958, + author = {Gleason, Andrew M.}, + title = {Projective topological spaces}, + journal = {Illinois J. Math.}, + fjournal = {Illinois Journal of Mathematics}, + volume = {2}, + year = {1958}, + pages = {482--489}, + issn = {0019-2082}, + mrclass = {54.00}, + mrnumber = {121775}, + mrreviewer = {Dana Scott}, + url = {http://projecteuclid.org/euclid.ijm/1255454110} +} + @Book{ goerss-jardine-2009, author = {Goerss, Paul G. and Jardine, John F.}, title = {Simplicial homotopy theory}, diff --git a/src/topology/extremally_disconnected.lean b/src/topology/extremally_disconnected.lean new file mode 100644 index 0000000000000..e09313a7f6680 --- /dev/null +++ b/src/topology/extremally_disconnected.lean @@ -0,0 +1,114 @@ +/- +Copyright (c) 2021 Johan Commelin. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johan Commelin +-/ +import topology.stone_cech + +/-! +# Extremally disconnected spaces + +An extremally disconnected topological space is a space in which the closure of every open set is +open. Such spaces are also called Stonean spaces. They are the projective objects in the category of +compact Hausdorff spaces. + +## Main declarations + +* `extremally_disconnected`: Predicate for a space to be extremally disconnected. +* `compact_t2.projective`: ¨Predicate for a topological space to be a projective object in the + category of compact Hausdorff spaces. +* `compact_t2.projective.extremally_disconnected`: Compact Hausdorff spaces that are + projective are extremally disconnected. + +# TODO + +Prove the converse to `compact_t2.projective.extremally_disconnected`, namely that a compact, +Hausdorff, extremally disconnected space is a projective object in the category of compact Hausdorff +spaces. + +## References + +[Gleason, *Projective topological spaces*][gleason1958] +-/ + +noncomputable theory + +open set +open_locale classical + +universes u v w +variables (X : Type u) [topological_space X] + +open function + +/-- An extremally disconnected topological space is a space +in which the closure of every open set is open. -/ +class extremally_disconnected : Prop := +(open_closure : ∀ U : set X, is_open U → is_open (closure U)) + +section + +include X + +/-- The assertion `compact_t2.projective` states that given continuous maps +`f : X → Z` and `g : Y → Z` with `g` surjective between `t_2`, compact topological spaces, +there exists a continuous lift `h : X → Y`, such that `f = g ∘ h`. -/ +def compact_t2.projective : Prop := +Π {Y Z : Type u} [topological_space Y] [topological_space Z], + by exactI Π [compact_space Y] [t2_space Y] [compact_space Z] [t2_space Z], + Π {f : X → Z} {g : Y → Z} (hf : continuous f) (hg : continuous g) (g_sur : surjective g), + ∃ h : X → Y, continuous h ∧ g ∘ h = f + +end + +variable {X} + +lemma stone_cech.projective [discrete_topology X] : compact_t2.projective (stone_cech X) := +begin + introsI Y Z _tsY _tsZ _csY _t2Y _csZ _csZ f g hf hg g_sur, + let s : Z → Y := λ z, classical.some $ g_sur z, + have hs : g ∘ s = id := funext (λ z, classical.some_spec (g_sur z)), + let t := s ∘ f ∘ stone_cech_unit, + have ht : continuous t := continuous_of_discrete_topology, + let h : stone_cech X → Y := stone_cech_extend ht, + have hh : continuous h := continuous_stone_cech_extend ht, + refine ⟨h, hh, dense_range_stone_cech_unit.equalizer (hg.comp hh) hf _⟩, + rw [comp.assoc, stone_cech_extend_extends ht, ← comp.assoc, hs, comp.left_id], +end + +protected lemma compact_t2.projective.extremally_disconnected [compact_space X] [t2_space X] + (h : compact_t2.projective X) : + extremally_disconnected X := +begin + refine { open_closure := λ U hU, _ }, + let Z₁ : set (X × bool) := Uᶜ ×ˢ {tt}, + let Z₂ : set (X × bool) := closure U ×ˢ {ff}, + let Z : set (X × bool) := Z₁ ∪ Z₂, + have hZ₁₂ : disjoint Z₁ Z₂ := disjoint_left.2 (λ x hx₁ hx₂, by cases hx₁.2.symm.trans hx₂.2), + have hZ₁ : is_closed Z₁ := hU.is_closed_compl.prod (t1_space.t1 _), + have hZ₂ : is_closed Z₂ := is_closed_closure.prod (t1_space.t1 ff), + have hZ : is_closed Z := hZ₁.union hZ₂, + let f : Z → X := prod.fst ∘ subtype.val, + have f_cont : continuous f := continuous_fst.comp continuous_subtype_val, + have f_sur : surjective f, + { intro x, + by_cases hx : x ∈ U, + { exact ⟨⟨(x, ff), or.inr ⟨subset_closure hx, set.mem_singleton _⟩⟩, rfl⟩ }, + { exact ⟨⟨(x, tt), or.inl ⟨hx, set.mem_singleton _⟩⟩, rfl⟩ } }, + haveI : compact_space Z := is_compact_iff_compact_space.mp hZ.is_compact, + obtain ⟨g, hg, g_sec⟩ := h continuous_id f_cont f_sur, + let φ := coe ∘ g, + have hφ : continuous φ := continuous_subtype_val.comp hg, + have hφ₁ : ∀ x, (φ x).1 = x := congr_fun g_sec, + suffices : closure U = φ ⁻¹' Z₂, + { rw [this, set.preimage_comp, ←is_closed_compl_iff, ←preimage_compl, + ←preimage_subtype_coe_eq_compl subset.rfl], + { exact hZ₁.preimage hφ }, + { rw [hZ₁₂.inter_eq, inter_empty] } }, + refine (closure_minimal _ $ hZ₂.preimage hφ).antisymm (λ x hx, _), + { rintro x hx, + have : φ x ∈ (Z₁ ∪ Z₂) := (g x).2, + simpa [hx, hφ₁] using this }, + { rw ←hφ₁ x, + exact hx.1 } +end From 09079525fd01b3dda35e96adaa08d2f943e1648c Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Sat, 22 Apr 2023 22:59:28 +0000 Subject: [PATCH 0865/1291] =?UTF-8?q?chore(analysis/normed/group/seminorm)?= =?UTF-8?q?:=20use=20coe=5Fnonneg=20instead=20of=20=E2=84=9D=E2=89=A50.pro?= =?UTF-8?q?p=20(#18833)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The change in `analysis.seminorm` has already been forward-ported in https://github.com/leanprover-community/mathlib4/pull/3484 to fix an error. Presumably the other changes will prevent the corresponding errors when we get to porting these files. This is needed in Lean 4 because `.prop` is about `Subtype.val` but `.coe_nonneg` is about `NNReal.toReal`. The two are defeq (and there is a simp lemma `NNReal.val_eq_coe` to convert between them), but not reducibly equal. In Lean 3 both are just about `coe`, so the change doesn't matter. Co-authored-by: Scott Morrison --- src/analysis/normed/group/seminorm.lean | 6 +++--- src/analysis/seminorm.lean | 2 +- 2 files changed, 4 insertions(+), 4 deletions(-) diff --git a/src/analysis/normed/group/seminorm.lean b/src/analysis/normed/group/seminorm.lean index 08ea6081b6a2f..2386f19137eb2 100644 --- a/src/analysis/normed/group/seminorm.lean +++ b/src/analysis/normed/group/seminorm.lean @@ -338,7 +338,7 @@ instance [has_smul R' ℝ] [has_smul R' ℝ≥0] [is_scalar_tower R' ℝ≥0 ℝ lemma smul_sup (r : R) (p q : add_group_seminorm E) : r • (p ⊔ q) = r • p ⊔ r • q := have real.smul_max : ∀ x y : ℝ, r • max x y = max (r • x) (r • y), from λ x y, by simpa only [←smul_eq_mul, ←nnreal.smul_def, smul_one_smul ℝ≥0 r (_ : ℝ)] - using mul_max_of_nonneg x y (r • 1 : ℝ≥0).prop, + using mul_max_of_nonneg x y (r • 1 : ℝ≥0).coe_nonneg, ext $ λ x, real.smul_max _ _ end add_group_seminorm @@ -460,7 +460,7 @@ lemma smul_apply (r : R) (p : group_seminorm E) (x : E) : (r • p) x = r • p lemma smul_sup (r : R) (p q : group_seminorm E) : r • (p ⊔ q) = r • p ⊔ r • q := have real.smul_max : ∀ x y : ℝ, r • max x y = max (r • x) (r • y), from λ x y, by simpa only [←smul_eq_mul, ←nnreal.smul_def, smul_one_smul ℝ≥0 r (_ : ℝ)] - using mul_max_of_nonneg x y (r • 1 : ℝ≥0).prop, + using mul_max_of_nonneg x y (r • 1 : ℝ≥0).coe_nonneg, ext $ λ x, real.smul_max _ _ end group_seminorm @@ -507,7 +507,7 @@ lemma smul_apply (r : R) (p : nonarch_add_group_seminorm E) (x : E) : (r • p) lemma smul_sup (r : R) (p q : nonarch_add_group_seminorm E) : r • (p ⊔ q) = r • p ⊔ r • q := have real.smul_max : ∀ x y : ℝ, r • max x y = max (r • x) (r • y), from λ x y, by simpa only [←smul_eq_mul, ←nnreal.smul_def, smul_one_smul ℝ≥0 r (_ : ℝ)] - using mul_max_of_nonneg x y (r • 1 : ℝ≥0).prop, + using mul_max_of_nonneg x y (r • 1 : ℝ≥0).coe_nonneg, ext $ λ x, real.smul_max _ _ end nonarch_add_group_seminorm diff --git a/src/analysis/seminorm.lean b/src/analysis/seminorm.lean index ded9d63911138..517b0787f6322 100644 --- a/src/analysis/seminorm.lean +++ b/src/analysis/seminorm.lean @@ -207,7 +207,7 @@ lemma smul_sup [has_smul R ℝ] [has_smul R ℝ≥0] [is_scalar_tower R ℝ≥0 r • (p ⊔ q) = r • p ⊔ r • q := have real.smul_max : ∀ x y : ℝ, r • max x y = max (r • x) (r • y), from λ x y, by simpa only [←smul_eq_mul, ←nnreal.smul_def, smul_one_smul ℝ≥0 r (_ : ℝ)] - using mul_max_of_nonneg x y (r • 1 : ℝ≥0).prop, + using mul_max_of_nonneg x y (r • 1 : ℝ≥0).coe_nonneg, ext $ λ x, real.smul_max _ _ instance : partial_order (seminorm 𝕜 E) := From 473ee9e095b89a24df125a4d363bd996f0989748 Mon Sep 17 00:00:00 2001 From: Eric Rodriguez Date: Sun, 23 Apr 2023 21:28:44 +0000 Subject: [PATCH 0866/1291] chore(zmod/basic): remove unneeded import (#18856) This file was importing a weird dependency; I am now minimising this. --- src/data/zmod/basic.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/data/zmod/basic.lean b/src/data/zmod/basic.lean index bf03ebdeae0e2..6fb43d1ba331c 100644 --- a/src/data/zmod/basic.lean +++ b/src/data/zmod/basic.lean @@ -6,8 +6,8 @@ Authors: Chris Hughes import algebra.char_p.basic import data.nat.parity -import algebra.group.conj_finite import tactic.fin_cases +import data.fintype.units /-! # Integers mod `n` From 74ad1c88c77e799d2fea62801d1dbbd698cff1b7 Mon Sep 17 00:00:00 2001 From: Eric Rodriguez <37984851+ericrbg@users.noreply.github.com> Date: Mon, 24 Apr 2023 00:07:33 +0000 Subject: [PATCH 0867/1291] chore(zmod/basic): alphabetise imports (#18858) This wasn't noticed in #18856 --- src/data/zmod/basic.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/data/zmod/basic.lean b/src/data/zmod/basic.lean index 6fb43d1ba331c..807cd3aff05e8 100644 --- a/src/data/zmod/basic.lean +++ b/src/data/zmod/basic.lean @@ -5,9 +5,9 @@ Authors: Chris Hughes -/ import algebra.char_p.basic +import data.fintype.units import data.nat.parity import tactic.fin_cases -import data.fintype.units /-! # Integers mod `n` From a07a7ae98384cd6485d7825e161e528ba78ef3bc Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 24 Apr 2023 09:08:24 +0000 Subject: [PATCH 0868/1291] refactor(linear_algebra/matrix/nonsingular_inverse): move results about block matrices to `schur_complement` (#18850) This gets these out of the critical path for porting, and also balances the size of these files a little better. I've added myself second rather than last on the author list since the bulk of these moved lemmas are mine, and they are about equal in number of lines to the existing lemmas. The lemmas have been moved without modification, though the module docstring for `schur_complement` has been adapted accordingly. --- .../matrix/nonsingular_inverse.lean | 89 ------------ .../matrix/schur_complement.lean | 132 +++++++++++++++--- 2 files changed, 113 insertions(+), 108 deletions(-) diff --git a/src/linear_algebra/matrix/nonsingular_inverse.lean b/src/linear_algebra/matrix/nonsingular_inverse.lean index 1bb03fe96276e..c730f21c63493 100644 --- a/src/linear_algebra/matrix/nonsingular_inverse.lean +++ b/src/linear_algebra/matrix/nonsingular_inverse.lean @@ -615,39 +615,6 @@ inv_submatrix_equiv A e₁.symm e₂.symm end submatrix -/-! ### Block matrices -/ - -section block -variables [fintype l] -variables [decidable_eq l] -variables [fintype m] -variables [decidable_eq m] - -/-- LDU decomposition of a block matrix with an invertible top-left corner, using the -Schur complement. -/ -lemma from_blocks_eq_of_invertible₁₁ - (A : matrix m m α) (B : matrix m n α) (C : matrix l m α) (D : matrix l n α) [invertible A] : - from_blocks A B C D = - from_blocks 1 0 (C⬝⅟A) 1 ⬝ from_blocks A 0 0 (D - C⬝(⅟A)⬝B) ⬝ from_blocks 1 (⅟A⬝B) 0 1 := -by simp only [from_blocks_multiply, matrix.mul_zero, matrix.zero_mul, add_zero, zero_add, - matrix.one_mul, matrix.mul_one, matrix.inv_of_mul_self, matrix.mul_inv_of_self_assoc, - matrix.mul_inv_of_mul_self_cancel, matrix.mul_assoc, add_sub_cancel'_right] - -/-- LDU decomposition of a block matrix with an invertible bottom-right corner, using the -Schur complement. -/ -lemma from_blocks_eq_of_invertible₂₂ - (A : matrix l m α) (B : matrix l n α) (C : matrix n m α) (D : matrix n n α) [invertible D] : - from_blocks A B C D = - from_blocks 1 (B⬝⅟D) 0 1 ⬝ from_blocks (A - B⬝⅟D⬝C) 0 0 D ⬝ from_blocks 1 0 (⅟D ⬝ C) 1 := -(matrix.reindex (equiv.sum_comm _ _) (equiv.sum_comm _ _)).injective $ by - simpa [reindex_apply, sum_comm_symm, - ←submatrix_mul_equiv _ _ _ (equiv.sum_comm n m), - ←submatrix_mul_equiv _ _ _ (equiv.sum_comm n l), - sum_comm_apply, from_blocks_submatrix_sum_swap_sum_swap] - using from_blocks_eq_of_invertible₁₁ D C B A - -end block - /-! ### More results about determinants -/ section det @@ -663,62 +630,6 @@ lemma det_conj' {M : matrix m m α} (h : is_unit M) (N : matrix m m α) : det (M⁻¹ ⬝ N ⬝ M) = det N := by rw [←h.unit_spec, ←coe_units_inv, det_units_conj'] -/-- Determinant of a 2×2 block matrix, expanded around an invertible top left element in terms of -the Schur complement. -/ -lemma det_from_blocks₁₁ (A : matrix m m α) (B : matrix m n α) (C : matrix n m α) (D : matrix n n α) - [invertible A] : (matrix.from_blocks A B C D).det = det A * det (D - C ⬝ (⅟A) ⬝ B) := -by rw [from_blocks_eq_of_invertible₁₁, det_mul, det_mul, det_from_blocks_zero₂₁, - det_from_blocks_zero₂₁, det_from_blocks_zero₁₂, det_one, det_one, one_mul, one_mul, mul_one] - -@[simp] lemma det_from_blocks_one₁₁ (B : matrix m n α) (C : matrix n m α) (D : matrix n n α) : - (matrix.from_blocks 1 B C D).det = det (D - C ⬝ B) := -begin - haveI : invertible (1 : matrix m m α) := invertible_one, - rw [det_from_blocks₁₁, inv_of_one, matrix.mul_one, det_one, one_mul], -end - -/-- Determinant of a 2×2 block matrix, expanded around an invertible bottom right element in terms -of the Schur complement. -/ -lemma det_from_blocks₂₂ (A : matrix m m α) (B : matrix m n α) (C : matrix n m α) (D : matrix n n α) - [invertible D] : (matrix.from_blocks A B C D).det = det D * det (A - B ⬝ (⅟D) ⬝ C) := -begin - have : from_blocks A B C D = (from_blocks D C B A).submatrix (sum_comm _ _) (sum_comm _ _), - { ext i j, - cases i; cases j; refl }, - rw [this, det_submatrix_equiv_self, det_from_blocks₁₁], -end - -@[simp] lemma det_from_blocks_one₂₂ (A : matrix m m α) (B : matrix m n α) (C : matrix n m α) : - (matrix.from_blocks A B C 1).det = det (A - B ⬝ C) := -begin - haveI : invertible (1 : matrix n n α) := invertible_one, - rw [det_from_blocks₂₂, inv_of_one, matrix.mul_one, det_one, one_mul], -end - -/-- The **Weinstein–Aronszajn identity**. Note the `1` on the LHS is of shape m×m, while the `1` on -the RHS is of shape n×n. -/ -lemma det_one_add_mul_comm (A : matrix m n α) (B : matrix n m α) : - det (1 + A ⬝ B) = det (1 + B ⬝ A) := -calc det (1 + A ⬝ B) - = det (from_blocks 1 (-A) B 1) : by rw [det_from_blocks_one₂₂, matrix.neg_mul, sub_neg_eq_add] -... = det (1 + B ⬝ A) : by rw [det_from_blocks_one₁₁, matrix.mul_neg, sub_neg_eq_add] - -/-- Alternate statement of the **Weinstein–Aronszajn identity** -/ -lemma det_mul_add_one_comm (A : matrix m n α) (B : matrix n m α) : - det (A ⬝ B + 1) = det (B ⬝ A + 1) := -by rw [add_comm, det_one_add_mul_comm, add_comm] - -lemma det_one_sub_mul_comm (A : matrix m n α) (B : matrix n m α) : - det (1 - A ⬝ B) = det (1 - B ⬝ A) := -by rw [sub_eq_add_neg, ←matrix.neg_mul, det_one_add_mul_comm, matrix.mul_neg, ←sub_eq_add_neg] - -/-- A special case of the **Matrix determinant lemma** for when `A = I`. - -TODO: show this more generally. -/ -lemma det_one_add_col_mul_row (u v : m → α) : det (1 + col u ⬝ row v) = 1 + v ⬝ᵥ u := -by rw [det_one_add_mul_comm, det_unique, pi.add_apply, pi.add_apply, matrix.one_apply_eq, - matrix.row_mul_col_apply] - end det end matrix diff --git a/src/linear_algebra/matrix/schur_complement.lean b/src/linear_algebra/matrix/schur_complement.lean index bda40a17b46a9..da1b1d46d3a3b 100644 --- a/src/linear_algebra/matrix/schur_complement.lean +++ b/src/linear_algebra/matrix/schur_complement.lean @@ -1,30 +1,130 @@ /- Copyright (c) 2022 Alexander Bentkamp. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Alexander Bentkamp, Jeremy Avigad, Johan Commelin +Authors: Alexander Bentkamp, Eric Wieser, Jeremy Avigad, Johan Commelin -/ import linear_algebra.matrix.nonsingular_inverse import linear_algebra.matrix.pos_def -/-! # Schur complement +/-! # 2×2 block matrices and the Schur complement -This file proves properties of the Schur complement `D - C A⁻¹ B` of a block matrix `[A B; C D]`. +This file proves properties of 2×2 block matrices `[A B; C D]` that relate to the Schur complement +`D - C⬝A⁻¹⬝B`. -The determinant of a block matrix in terms of the Schur complement is expressed in the lemmas -`matrix.det_from_blocks₁₁` and `matrix.det_from_blocks₂₂` in the file -`linear_algebra.matrix.nonsingular_inverse`. +## Main results -## Main result - - * `matrix.schur_complement_pos_semidef_iff` : If a matrix `A` is positive definite, then `[A B; Bᴴ - D]` is postive semidefinite if and only if `D - Bᴴ A⁻¹ B` is postive semidefinite. + * `matrix.det_from_blocks₁₁`, `matrix.det_from_blocks₂₂`: determinant of a block matrix in terms of + the Schur complement. + * `matrix.det_one_add_mul_comm`: the **Weinstein–Aronszajn identity**. + * `matrix.schur_complement_pos_semidef_iff` : If a matrix `A` is positive definite, then + `[A B; Bᴴ D]` is postive semidefinite if and only if `D - Bᴴ A⁻¹ B` is postive semidefinite. -/ +variables {l m n α : Type*} + namespace matrix +open_locale matrix + +section comm_ring +variables [fintype l] [fintype m] [fintype n] +variables [decidable_eq l] [decidable_eq m] [decidable_eq n] +variables [comm_ring α] + +/-- LDU decomposition of a block matrix with an invertible top-left corner, using the +Schur complement. -/ +lemma from_blocks_eq_of_invertible₁₁ + (A : matrix m m α) (B : matrix m n α) (C : matrix l m α) (D : matrix l n α) [invertible A] : + from_blocks A B C D = + from_blocks 1 0 (C⬝⅟A) 1 ⬝ from_blocks A 0 0 (D - C⬝(⅟A)⬝B) ⬝ from_blocks 1 (⅟A⬝B) 0 1 := +by simp only [from_blocks_multiply, matrix.mul_zero, matrix.zero_mul, add_zero, zero_add, + matrix.one_mul, matrix.mul_one, matrix.inv_of_mul_self, matrix.mul_inv_of_self_assoc, + matrix.mul_inv_of_mul_self_cancel, matrix.mul_assoc, add_sub_cancel'_right] + +/-- LDU decomposition of a block matrix with an invertible bottom-right corner, using the +Schur complement. -/ +lemma from_blocks_eq_of_invertible₂₂ + (A : matrix l m α) (B : matrix l n α) (C : matrix n m α) (D : matrix n n α) [invertible D] : + from_blocks A B C D = + from_blocks 1 (B⬝⅟D) 0 1 ⬝ from_blocks (A - B⬝⅟D⬝C) 0 0 D ⬝ from_blocks 1 0 (⅟D ⬝ C) 1 := +(matrix.reindex (equiv.sum_comm _ _) (equiv.sum_comm _ _)).injective $ by + simpa [reindex_apply, equiv.sum_comm_symm, + ←submatrix_mul_equiv _ _ _ (equiv.sum_comm n m), + ←submatrix_mul_equiv _ _ _ (equiv.sum_comm n l), + equiv.sum_comm_apply, from_blocks_submatrix_sum_swap_sum_swap] + using from_blocks_eq_of_invertible₁₁ D C B A + +/-! ### Lemmas about `matrix.det` -/ + +section det + +/-- Determinant of a 2×2 block matrix, expanded around an invertible top left element in terms of +the Schur complement. -/ +lemma det_from_blocks₁₁ (A : matrix m m α) (B : matrix m n α) (C : matrix n m α) (D : matrix n n α) + [invertible A] : (matrix.from_blocks A B C D).det = det A * det (D - C ⬝ (⅟A) ⬝ B) := +by rw [from_blocks_eq_of_invertible₁₁, det_mul, det_mul, det_from_blocks_zero₂₁, + det_from_blocks_zero₂₁, det_from_blocks_zero₁₂, det_one, det_one, one_mul, one_mul, mul_one] + +@[simp] lemma det_from_blocks_one₁₁ (B : matrix m n α) (C : matrix n m α) (D : matrix n n α) : + (matrix.from_blocks 1 B C D).det = det (D - C ⬝ B) := +begin + haveI : invertible (1 : matrix m m α) := invertible_one, + rw [det_from_blocks₁₁, inv_of_one, matrix.mul_one, det_one, one_mul], +end + +/-- Determinant of a 2×2 block matrix, expanded around an invertible bottom right element in terms +of the Schur complement. -/ +lemma det_from_blocks₂₂ (A : matrix m m α) (B : matrix m n α) (C : matrix n m α) (D : matrix n n α) + [invertible D] : (matrix.from_blocks A B C D).det = det D * det (A - B ⬝ (⅟D) ⬝ C) := +begin + have : from_blocks A B C D + = (from_blocks D C B A).submatrix (equiv.sum_comm _ _) (equiv.sum_comm _ _), + { ext i j, + cases i; cases j; refl }, + rw [this, det_submatrix_equiv_self, det_from_blocks₁₁], +end + +@[simp] lemma det_from_blocks_one₂₂ (A : matrix m m α) (B : matrix m n α) (C : matrix n m α) : + (matrix.from_blocks A B C 1).det = det (A - B ⬝ C) := +begin + haveI : invertible (1 : matrix n n α) := invertible_one, + rw [det_from_blocks₂₂, inv_of_one, matrix.mul_one, det_one, one_mul], +end + +/-- The **Weinstein–Aronszajn identity**. Note the `1` on the LHS is of shape m×m, while the `1` on +the RHS is of shape n×n. -/ +lemma det_one_add_mul_comm (A : matrix m n α) (B : matrix n m α) : + det (1 + A ⬝ B) = det (1 + B ⬝ A) := +calc det (1 + A ⬝ B) + = det (from_blocks 1 (-A) B 1) : by rw [det_from_blocks_one₂₂, matrix.neg_mul, sub_neg_eq_add] +... = det (1 + B ⬝ A) : by rw [det_from_blocks_one₁₁, matrix.mul_neg, sub_neg_eq_add] + +/-- Alternate statement of the **Weinstein–Aronszajn identity** -/ +lemma det_mul_add_one_comm (A : matrix m n α) (B : matrix n m α) : + det (A ⬝ B + 1) = det (B ⬝ A + 1) := +by rw [add_comm, det_one_add_mul_comm, add_comm] + +lemma det_one_sub_mul_comm (A : matrix m n α) (B : matrix n m α) : + det (1 - A ⬝ B) = det (1 - B ⬝ A) := +by rw [sub_eq_add_neg, ←matrix.neg_mul, det_one_add_mul_comm, matrix.mul_neg, ←sub_eq_add_neg] + +/-- A special case of the **Matrix determinant lemma** for when `A = I`. + +TODO: show this more generally. -/ +lemma det_one_add_col_mul_row (u v : m → α) : det (1 + col u ⬝ row v) = 1 + v ⬝ᵥ u := +by rw [det_one_add_mul_comm, det_unique, pi.add_apply, pi.add_apply, matrix.one_apply_eq, + matrix.row_mul_col_apply] + +end det + +end comm_ring + +/-! ### Lemmas about `ℝ` and `ℂ`-/ + +section is_R_or_C open_locale matrix -variables {n : Type*} {m : Type*} {𝕜 : Type*} [is_R_or_C 𝕜] +variables {𝕜 : Type*} [is_R_or_C 𝕜] localized "infix ` ⊕ᵥ `:65 := sum.elim" in matrix @@ -54,14 +154,6 @@ begin abel end -end matrix - -namespace matrix - -open_locale matrix -variables {n : Type*} {m : Type*} - {𝕜 : Type*} [is_R_or_C 𝕜] - lemma is_hermitian.from_blocks₁₁ [fintype m] [decidable_eq m] {A : matrix m m 𝕜} (B : matrix m n 𝕜) (D : matrix n n 𝕜) (hA : A.is_hermitian) : @@ -121,4 +213,6 @@ begin convert pos_semidef.from_blocks₁₁ _ _ hD; apply_instance <|> simp end +end is_R_or_C + end matrix From 2651125b48fc5c170ab1111afd0817c903b1fc6c Mon Sep 17 00:00:00 2001 From: Eric Rodriguez Date: Mon, 24 Apr 2023 13:09:26 +0000 Subject: [PATCH 0869/1291] feat(algebra & polynomial): some (q)smul lemmas+generalisations (#18852) There is many generalisations around these areas too, but I am specifically not doing them as it will be easier done after the port. I am only doing what I need for merging in the splitting field diamond fix. Co-authored-by: Eric Rodriguez <37984851+ericrbg@users.noreply.github.com> --- src/algebra/algebra/basic.lean | 6 ------ src/algebra/field/defs.lean | 4 ++++ src/algebra/monoid_algebra/basic.lean | 3 +++ src/data/finsupp/basic.lean | 24 +++++++++++------------- src/data/polynomial/basic.lean | 17 +++++++++++++++-- src/data/polynomial/coeff.lean | 2 +- 6 files changed, 34 insertions(+), 22 deletions(-) diff --git a/src/algebra/algebra/basic.lean b/src/algebra/algebra/basic.lean index 431f754bc9471..261a01004e27b 100644 --- a/src/algebra/algebra/basic.lean +++ b/src/algebra/algebra/basic.lean @@ -606,12 +606,6 @@ by rw [←algebra.commutes, ←algebra.commutes, map_algebra_map_mul] end linear_map - -@[simp] lemma rat.smul_one_eq_coe {A : Type*} [division_ring A] [algebra ℚ A] (m : ℚ) : - @@has_smul.smul algebra.to_has_smul m (1 : A) = ↑m := -by rw [algebra.smul_def, mul_one, eq_rat_cast] - - section nat variables {R : Type*} [semiring R] diff --git a/src/algebra/field/defs.lean b/src/algebra/field/defs.lean index 098aabbec6949..73ddad343d0d0 100644 --- a/src/algebra/field/defs.lean +++ b/src/algebra/field/defs.lean @@ -143,6 +143,10 @@ instance smul_division_ring : has_smul ℚ K := lemma smul_def (a : ℚ) (x : K) : a • x = ↑a * x := division_ring.qsmul_eq_mul' a x +@[simp] lemma smul_one_eq_coe (A : Type*) [division_ring A] (m : ℚ) : + m • (1 : A) = ↑m := +by rw [rat.smul_def, mul_one] + end rat end division_ring diff --git a/src/algebra/monoid_algebra/basic.lean b/src/algebra/monoid_algebra/basic.lean index f746821df66ff..53dce47b0586b 100644 --- a/src/algebra/monoid_algebra/basic.lean +++ b/src/algebra/monoid_algebra/basic.lean @@ -1123,6 +1123,9 @@ instance [comm_ring k] [add_comm_monoid G] : comm_ring (add_monoid_algebra k G) variables {S : Type*} +instance [semiring k] [distrib_smul R k] : distrib_smul R (add_monoid_algebra k G) := +finsupp.distrib_smul G k + instance [monoid R] [semiring k] [distrib_mul_action R k] : distrib_mul_action R (add_monoid_algebra k G) := finsupp.distrib_mul_action G k diff --git a/src/data/finsupp/basic.lean b/src/data/finsupp/basic.lean index 4b1e6bda95cf5..de49e4b2009a1 100644 --- a/src/data/finsupp/basic.lean +++ b/src/data/finsupp/basic.lean @@ -1259,16 +1259,16 @@ Throughout this section, some `monoid` and `semiring` arguments are specified wi `[]`. See note [implicit instance arguments]. -/ -@[simp] lemma coe_smul [add_monoid M] [distrib_smul R M] +@[simp] lemma coe_smul [add_monoid M] [smul_zero_class R M] (b : R) (v : α →₀ M) : ⇑(b • v) = b • v := rfl -lemma smul_apply [add_monoid M] [distrib_smul R M] +lemma smul_apply [add_monoid M] [smul_zero_class R M] (b : R) (v : α →₀ M) (a : α) : (b • v) a = b • (v a) := rfl -lemma _root_.is_smul_regular.finsupp [add_monoid M] [distrib_smul R M] {k : R} +lemma _root_.is_smul_regular.finsupp [add_monoid M] [smul_zero_class R M] {k : R} (hk : is_smul_regular M k) : is_smul_regular (α →₀ M) k := λ _ _ h, ext $ λ i, hk (congr_fun h i) -instance [nonempty α] [add_monoid M] [distrib_smul R M] [has_faithful_smul R M] : +instance [nonempty α] [add_monoid M] [smul_zero_class R M] [has_faithful_smul R M] : has_faithful_smul R (α →₀ M) := { eq_of_smul_eq_smul := λ r₁ r₂ h, let ⟨a⟩ := ‹nonempty α› in eq_of_smul_eq_smul $ λ m : M, by simpa using congr_fun (h (single a m)) a } @@ -1286,18 +1286,16 @@ instance [monoid R] [add_monoid M] [distrib_mul_action R M] : distrib_mul_action mul_smul := λ r s x, ext $ λ _, mul_smul _ _ _, ..finsupp.distrib_smul _ _ } -instance [monoid R] [monoid S] [add_monoid M] [distrib_mul_action R M] [distrib_mul_action S M] - [has_smul R S] [is_scalar_tower R S M] : - is_scalar_tower R S (α →₀ M) := +instance [has_zero M] [smul_zero_class R M] [smul_zero_class S M] [has_smul R S] + [is_scalar_tower R S M] : is_scalar_tower R S (α →₀ M) := { smul_assoc := λ r s a, ext $ λ _, smul_assoc _ _ _ } -instance [monoid R] [monoid S] [add_monoid M] [distrib_mul_action R M] [distrib_mul_action S M] - [smul_comm_class R S M] : - smul_comm_class R S (α →₀ M) := +instance [has_zero M] [smul_zero_class R M] [smul_zero_class S M] + [smul_comm_class R S M] : smul_comm_class R S (α →₀ M) := { smul_comm := λ r s a, ext $ λ _, smul_comm _ _ _ } -instance [monoid R] [add_monoid M] [distrib_mul_action R M] [distrib_mul_action Rᵐᵒᵖ M] - [is_central_scalar R M] : is_central_scalar R (α →₀ M) := +instance [has_zero M] [smul_zero_class R M] [smul_zero_class Rᵐᵒᵖ M] [is_central_scalar R M] : + is_central_scalar R (α →₀ M) := { op_smul_eq_smul := λ r a, ext $ λ _, op_smul_eq_smul _ _ } instance [semiring R] [add_comm_monoid M] [module R M] : module R (α →₀ M) := @@ -1332,7 +1330,7 @@ lemma map_domain_smul {_ : monoid R} [add_comm_monoid M] [distrib_mul_action R M {f : α → β} (b : R) (v : α →₀ M) : map_domain f (b • v) = b • map_domain f v := map_domain_map_range _ _ _ _ (smul_add b) -@[simp] lemma smul_single {_ : monoid R} [add_monoid M] [distrib_mul_action R M] +@[simp] lemma smul_single [has_zero M] [smul_zero_class R M] (c : R) (a : α) (b : M) : c • finsupp.single a b = finsupp.single a (c • b) := map_range_single diff --git a/src/data/polynomial/basic.lean b/src/data/polynomial/basic.lean index 7334c8cdcc10c..5bd63b93712af 100644 --- a/src/data/polynomial/basic.lean +++ b/src/data/polynomial/basic.lean @@ -183,6 +183,10 @@ function.injective.semiring to_finsupp to_finsupp_injective to_finsupp_zero to_finsupp_one to_finsupp_add to_finsupp_mul (λ _ _, to_finsupp_smul _ _) to_finsupp_pow (λ _, rfl) +instance {S} [distrib_smul S R] : distrib_smul S R[X] := +function.injective.distrib_smul ⟨to_finsupp, to_finsupp_zero, to_finsupp_add⟩ +to_finsupp_injective to_finsupp_smul + instance {S} [monoid S] [distrib_mul_action S R] : distrib_mul_action S R[X] := function.injective.distrib_mul_action ⟨to_finsupp, to_finsupp_zero, to_finsupp_add⟩ to_finsupp_injective to_finsupp_smul @@ -301,7 +305,7 @@ begin { simp [pow_succ, ih, monomial_mul_monomial, nat.succ_eq_add_one, mul_add, add_comm] }, end -lemma smul_monomial {S} [monoid S] [distrib_mul_action S R] (a : S) (n : ℕ) (b : R) : +lemma smul_monomial {S} [smul_zero_class S R] (a : S) (n : ℕ) (b : R) : a • monomial n b = monomial n (a • b) := to_finsupp_injective $ by simp @@ -342,7 +346,7 @@ lemma C_mul : C (a * b) = C a * C b := C.map_mul a b lemma C_add : C (a + b) = C a + C b := C.map_add a b -@[simp] lemma smul_C {S} [monoid S] [distrib_mul_action S R] (s : S) (r : R) : +@[simp] lemma smul_C {S} [smul_zero_class S R] (s : S) (r : R) : s • C r = C (s • r) := smul_monomial _ _ r @@ -857,6 +861,15 @@ mt (congr_arg (λ p, coeff p 1)) (by simp) end nonzero_semiring +section division_ring + +variables [division_ring R] + +lemma rat_smul_eq_C_mul (a : ℚ) (f : R[X]) : a • f = polynomial.C ↑a * f := +by rw [←rat.smul_one_eq_coe, ←polynomial.smul_C, C_1, smul_one_mul] + +end division_ring + @[simp] lemma nontrivial_iff [semiring R] : nontrivial R[X] ↔ nontrivial R := ⟨λ h, let ⟨r, s, hrs⟩ := @exists_pair_ne _ h in nontrivial.of_polynomial_ne hrs, λ h, @polynomial.nontrivial _ _ h⟩ diff --git a/src/data/polynomial/coeff.lean b/src/data/polynomial/coeff.lean index 8bcb8e2638580..517a60b716ae3 100644 --- a/src/data/polynomial/coeff.lean +++ b/src/data/polynomial/coeff.lean @@ -41,7 +41,7 @@ by { rcases p, rcases q, simp_rw [←of_finsupp_add, coeff], exact finsupp.add_a @[simp] lemma coeff_bit0 (p : R[X]) (n : ℕ) : coeff (bit0 p) n = bit0 (coeff p n) := by simp [bit0] -@[simp] lemma coeff_smul [monoid S] [distrib_mul_action S R] (r : S) (p : R[X]) (n : ℕ) : +@[simp] lemma coeff_smul [smul_zero_class S R] (r : S) (p : R[X]) (n : ℕ) : coeff (r • p) n = r • coeff p n := by { rcases p, simp_rw [←of_finsupp_smul, coeff], exact finsupp.smul_apply _ _ _ } From aa68866e61a7f1f40e7d6c8b0116ebd6370525c3 Mon Sep 17 00:00:00 2001 From: sgouezel Date: Tue, 25 Apr 2023 09:29:58 +0000 Subject: [PATCH 0870/1291] feat(analysis/calculus/cont_diff): bound for the n-th derivative of a composition (#18643) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We prove the bound `n! * C * D^n` for the `n`-th derivative of `g ∘ f` if the derivatives of `g` are bounded by `C` and the `i`-th derivative of `f` is bounded by `D^i`. --- src/analysis/calculus/cont_diff.lean | 265 ++++++++++++++++++++++++--- 1 file changed, 239 insertions(+), 26 deletions(-) diff --git a/src/analysis/calculus/cont_diff.lean b/src/analysis/calculus/cont_diff.lean index 790054897f53b..45b377d3401ec 100644 --- a/src/analysis/calculus/cont_diff.lean +++ b/src/analysis/calculus/cont_diff.lean @@ -6,12 +6,22 @@ Authors: Sébastien Gouëzel, Floris van Doorn import analysis.calculus.cont_diff_def import analysis.calculus.mean_value import analysis.normed_space.finite_dimension +import data.nat.choose.cast /-! # Higher differentiability of usual operations We prove that the usual operations (addition, multiplication, difference, composition, and -so on) preserve `C^n` functions. We also expand the API aound `C^n` functions. +so on) preserve `C^n` functions. We also expand the API around `C^n` functions. + +## Main results + +* `cont_diff.comp` states that the composition of two `C^n` functions is `C^n`. +* `norm_iterated_fderiv_comp_le` gives the bound `n! * C * D ^ n` for the `n`-th derivative + of `g ∘ f` assuming that the derivatives of `g` are bounded by `C` and the `i`-th + derivative of `f` is bounded by `D ^ i`. + +Similar results are given for `C^n` functions on domains. ## Notations @@ -26,7 +36,7 @@ derivative, differentiability, higher derivative, `C^n`, multilinear, Taylor ser -/ noncomputable theory -open_locale classical big_operators nnreal +open_locale classical big_operators nnreal nat local notation `∞` := (⊤ : ℕ∞) @@ -284,7 +294,7 @@ cont_diff_on_univ.1 $ cont_diff_on.continuous_linear_map_comp obtained by applying the linear map to the iterated derivative. -/ lemma continuous_linear_map.iterated_fderiv_within_comp_left {f : E → F} (g : F →L[𝕜] G) (hf : cont_diff_on 𝕜 n f s) (hs : unique_diff_on 𝕜 s) (hx : x ∈ s) - {i : ℕ} (hi : (i : with_top ℕ) ≤ n) : + {i : ℕ} (hi : (i : ℕ∞) ≤ n) : iterated_fderiv_within 𝕜 i (g ∘ f) s x = g.comp_continuous_multilinear_map (iterated_fderiv_within 𝕜 i f s x) := (((hf.ftaylor_series_within hs).continuous_linear_map_comp g).eq_ftaylor_series_of_unique_diff_on @@ -293,7 +303,7 @@ lemma continuous_linear_map.iterated_fderiv_within_comp_left /-- The iterated derivative of the composition with a linear map on the left is obtained by applying the linear map to the iterated derivative. -/ lemma continuous_linear_map.iterated_fderiv_comp_left - {f : E → F} (g : F →L[𝕜] G) (hf : cont_diff 𝕜 n f) (x : E) {i : ℕ} (hi : (i : with_top ℕ) ≤ n) : + {f : E → F} (g : F →L[𝕜] G) (hf : cont_diff 𝕜 n f) (x : E) {i : ℕ} (hi : (i : ℕ∞) ≤ n) : iterated_fderiv 𝕜 i (g ∘ f) x = g.comp_continuous_multilinear_map (iterated_fderiv 𝕜 i f x) := begin simp only [← iterated_fderiv_within_univ], @@ -330,7 +340,7 @@ end derivative within a set. -/ lemma linear_isometry.norm_iterated_fderiv_within_comp_left {f : E → F} (g : F →ₗᵢ[𝕜] G) (hf : cont_diff_on 𝕜 n f s) (hs : unique_diff_on 𝕜 s) (hx : x ∈ s) - {i : ℕ} (hi : (i : with_top ℕ) ≤ n) : + {i : ℕ} (hi : (i : ℕ∞) ≤ n) : ‖iterated_fderiv_within 𝕜 i (g ∘ f) s x‖ = ‖iterated_fderiv_within 𝕜 i f s x‖ := begin have : iterated_fderiv_within 𝕜 i (g ∘ f) s x = @@ -343,7 +353,7 @@ end /-- Composition with a linear isometry on the left preserves the norm of the iterated derivative. -/ lemma linear_isometry.norm_iterated_fderiv_comp_left - {f : E → F} (g : F →ₗᵢ[𝕜] G) (hf : cont_diff 𝕜 n f) (x : E) {i : ℕ} (hi : (i : with_top ℕ) ≤ n) : + {f : E → F} (g : F →ₗᵢ[𝕜] G) (hf : cont_diff 𝕜 n f) (x : E) {i : ℕ} (hi : (i : ℕ∞) ≤ n) : ‖iterated_fderiv 𝕜 i (g ∘ f) x‖ = ‖iterated_fderiv 𝕜 i f x‖ := begin simp only [← iterated_fderiv_within_univ], @@ -462,7 +472,7 @@ obtained by composing the iterated derivative with the linear map. -/ lemma continuous_linear_map.iterated_fderiv_within_comp_right {f : E → F} (g : G →L[𝕜] E) (hf : cont_diff_on 𝕜 n f s) (hs : unique_diff_on 𝕜 s) (h's : unique_diff_on 𝕜 (g⁻¹' s)) {x : G} - (hx : g x ∈ s) {i : ℕ} (hi : (i : with_top ℕ) ≤ n) : + (hx : g x ∈ s) {i : ℕ} (hi : (i : ℕ∞) ≤ n) : iterated_fderiv_within 𝕜 i (f ∘ g) (g ⁻¹' s) x = (iterated_fderiv_within 𝕜 i f s (g x)).comp_continuous_linear_map (λ _, g) := (((hf.ftaylor_series_within hs).comp_continuous_linear_map g).eq_ftaylor_series_of_unique_diff_on @@ -498,7 +508,7 @@ end /-- The iterated derivative of the composition with a linear map on the right is obtained by composing the iterated derivative with the linear map. -/ lemma continuous_linear_map.iterated_fderiv_comp_right - (g : G →L[𝕜] E) {f : E → F} (hf : cont_diff 𝕜 n f) (x : G) {i : ℕ} (hi : (i : with_top ℕ) ≤ n) : + (g : G →L[𝕜] E) {f : E → F} (hf : cont_diff 𝕜 n f) (x : G) {i : ℕ} (hi : (i : ℕ∞) ≤ n) : iterated_fderiv 𝕜 i (f ∘ g) x = (iterated_fderiv 𝕜 i f (g x)).comp_continuous_linear_map (λ _, g) := begin @@ -2368,7 +2378,7 @@ begin apply ((B (f x)).le_op_norm (g x)).trans, apply mul_le_mul_of_nonneg_right _ (norm_nonneg _), exact B.le_op_norm (f x) }, - { have In : (n : with_top ℕ) + 1 ≤ n.succ, by simp only [nat.cast_succ, le_refl], + { have In : (n : ℕ∞) + 1 ≤ n.succ, by simp only [nat.cast_succ, le_refl], have I1 : ‖iterated_fderiv_within 𝕜 n (λ (y : Du), B.precompR Du (f y) (fderiv_within 𝕜 g s y)) s x‖ ≤ ‖B‖ * ∑ (i : ℕ) in finset.range (n + 1), n.choose i * @@ -2413,7 +2423,7 @@ begin = iterated_fderiv_within 𝕜 n (λ y, B.precompR Du (f y) (fderiv_within 𝕜 g s y) + B.precompL Du (fderiv_within 𝕜 f s y) (g y)) s x, { apply iterated_fderiv_within_congr hs (λ y hy, _) hx, - have L : (1 : with_top ℕ) ≤ n.succ, + have L : (1 : ℕ∞) ≤ n.succ, by simpa only [enat.coe_one, nat.one_le_cast] using nat.succ_pos n, exact B.fderiv_within_of_bilinear (hf.differentiable_on L y hy) (hg.differentiable_on L y hy) (hs y hy) }, @@ -2436,9 +2446,9 @@ end iterated derivatives of `f` and `g` when `B` is bilinear: `‖D^n (x ↦ B (f x) (g x))‖ ≤ ‖B‖ ∑_{k ≤ n} n.choose k ‖D^k f‖ ‖D^{n-k} g‖` -/ lemma continuous_linear_map.norm_iterated_fderiv_within_le_of_bilinear - (B : E →L[𝕜] F →L[𝕜] G) {f : D → E} {g : D → F} {N : with_top ℕ} {s : set D} {x : D} + (B : E →L[𝕜] F →L[𝕜] G) {f : D → E} {g : D → F} {N : ℕ∞} {s : set D} {x : D} (hf : cont_diff_on 𝕜 N f s) (hg : cont_diff_on 𝕜 N g s) (hs : unique_diff_on 𝕜 s) (hx : x ∈ s) - {n : ℕ} (hn : (n : with_top ℕ) ≤ N) : + {n : ℕ} (hn : (n : ℕ∞) ≤ N) : ‖iterated_fderiv_within 𝕜 n (λ y, B (f y) (g y)) s x‖ ≤ ‖B‖ * ∑ i in finset.range (n+1), (n.choose i : ℝ) * ‖iterated_fderiv_within 𝕜 i f s x‖ * ‖iterated_fderiv_within 𝕜 (n-i) g s x‖ := @@ -2520,9 +2530,9 @@ end iterated derivatives of `f` and `g` when `B` is bilinear: `‖D^n (x ↦ B (f x) (g x))‖ ≤ ‖B‖ ∑_{k ≤ n} n.choose k ‖D^k f‖ ‖D^{n-k} g‖` -/ lemma continuous_linear_map.norm_iterated_fderiv_le_of_bilinear - (B : E →L[𝕜] F →L[𝕜] G) {f : D → E} {g : D → F} {N : with_top ℕ} + (B : E →L[𝕜] F →L[𝕜] G) {f : D → E} {g : D → F} {N : ℕ∞} (hf : cont_diff 𝕜 N f) (hg : cont_diff 𝕜 N g) (x : D) - {n : ℕ} (hn : (n : with_top ℕ) ≤ N) : + {n : ℕ} (hn : (n : ℕ∞) ≤ N) : ‖iterated_fderiv 𝕜 n (λ y, B (f y) (g y)) x‖ ≤ ‖B‖ * ∑ i in finset.range (n+1), (n.choose i : ℝ) * ‖iterated_fderiv 𝕜 i f x‖ * ‖iterated_fderiv 𝕜 (n-i) g x‖ := @@ -2536,9 +2546,9 @@ end iterated derivatives of `f` and `g` when `B` is bilinear of norm at most `1`: `‖D^n (x ↦ B (f x) (g x))‖ ≤ ∑_{k ≤ n} n.choose k ‖D^k f‖ ‖D^{n-k} g‖` -/ lemma continuous_linear_map.norm_iterated_fderiv_within_le_of_bilinear_of_le_one - (B : E →L[𝕜] F →L[𝕜] G) {f : D → E} {g : D → F} {N : with_top ℕ} {s : set D} {x : D} + (B : E →L[𝕜] F →L[𝕜] G) {f : D → E} {g : D → F} {N : ℕ∞} {s : set D} {x : D} (hf : cont_diff_on 𝕜 N f s) (hg : cont_diff_on 𝕜 N g s) (hs : unique_diff_on 𝕜 s) (hx : x ∈ s) - {n : ℕ} (hn : (n : with_top ℕ) ≤ N) (hB : ‖B‖ ≤ 1) : + {n : ℕ} (hn : (n : ℕ∞) ≤ N) (hB : ‖B‖ ≤ 1) : ‖iterated_fderiv_within 𝕜 n (λ y, B (f y) (g y)) s x‖ ≤ ∑ i in finset.range (n+1), (n.choose i : ℝ) * ‖iterated_fderiv_within 𝕜 i f s x‖ * ‖iterated_fderiv_within 𝕜 (n-i) g s x‖ := @@ -2552,9 +2562,9 @@ end iterated derivatives of `f` and `g` when `B` is bilinear of norm at most `1`: `‖D^n (x ↦ B (f x) (g x))‖ ≤ ∑_{k ≤ n} n.choose k ‖D^k f‖ ‖D^{n-k} g‖` -/ lemma continuous_linear_map.norm_iterated_fderiv_le_of_bilinear_of_le_one - (B : E →L[𝕜] F →L[𝕜] G) {f : D → E} {g : D → F} {N : with_top ℕ} + (B : E →L[𝕜] F →L[𝕜] G) {f : D → E} {g : D → F} {N : ℕ∞} (hf : cont_diff 𝕜 N f) (hg : cont_diff 𝕜 N g) (x : D) - {n : ℕ} (hn : (n : with_top ℕ) ≤ N) (hB : ‖B‖ ≤ 1) : + {n : ℕ} (hn : (n : ℕ∞) ≤ N) (hB : ‖B‖ ≤ 1) : ‖iterated_fderiv 𝕜 n (λ y, B (f y) (g y)) x‖ ≤ ∑ i in finset.range (n+1), (n.choose i : ℝ) * ‖iterated_fderiv 𝕜 i f x‖ * ‖iterated_fderiv 𝕜 (n-i) g x‖ := @@ -2570,8 +2580,8 @@ variables {𝕜' : Type*} [normed_field 𝕜'] [normed_algebra 𝕜 𝕜'] [norm [is_scalar_tower 𝕜 𝕜' F] lemma norm_iterated_fderiv_within_smul_le - {f : E → 𝕜'} {g : E → F} {N : with_top ℕ} (hf : cont_diff_on 𝕜 N f s) (hg : cont_diff_on 𝕜 N g s) - (hs : unique_diff_on 𝕜 s) {x : E} (hx : x ∈ s) {n : ℕ} (hn : (n : with_top ℕ) ≤ N) : + {f : E → 𝕜'} {g : E → F} {N : ℕ∞} (hf : cont_diff_on 𝕜 N f s) (hg : cont_diff_on 𝕜 N g s) + (hs : unique_diff_on 𝕜 s) {x : E} (hx : x ∈ s) {n : ℕ} (hn : (n : ℕ∞) ≤ N) : ‖iterated_fderiv_within 𝕜 n (λ y, f y • g y) s x‖ ≤ ∑ i in finset.range (n+1), (n.choose i : ℝ) * ‖iterated_fderiv_within 𝕜 i f s x‖ * ‖iterated_fderiv_within 𝕜 (n-i) g s x‖ := @@ -2580,8 +2590,8 @@ lemma norm_iterated_fderiv_within_smul_le continuous_linear_map.op_norm_lsmul_le lemma norm_iterated_fderiv_smul_le - {f : E → 𝕜'} {g : E → F} {N : with_top ℕ} (hf : cont_diff 𝕜 N f) (hg : cont_diff 𝕜 N g) - (x : E) {n : ℕ} (hn : (n : with_top ℕ) ≤ N) : + {f : E → 𝕜'} {g : E → F} {N : ℕ∞} (hf : cont_diff 𝕜 N f) (hg : cont_diff 𝕜 N g) + (x : E) {n : ℕ} (hn : (n : ℕ∞) ≤ N) : ‖iterated_fderiv 𝕜 n (λ y, f y • g y) x‖ ≤ ∑ i in finset.range (n+1), (n.choose i : ℝ) * ‖iterated_fderiv 𝕜 i f x‖ * ‖iterated_fderiv 𝕜 (n-i) g x‖ := @@ -2595,8 +2605,8 @@ section variables {A : Type*} [normed_ring A] [normed_algebra 𝕜 A] lemma norm_iterated_fderiv_within_mul_le - {f : E → A} {g : E → A} {N : with_top ℕ} (hf : cont_diff_on 𝕜 N f s) (hg : cont_diff_on 𝕜 N g s) - (hs : unique_diff_on 𝕜 s) {x : E} (hx : x ∈ s) {n : ℕ} (hn : (n : with_top ℕ) ≤ N) : + {f : E → A} {g : E → A} {N : ℕ∞} (hf : cont_diff_on 𝕜 N f s) (hg : cont_diff_on 𝕜 N g s) + (hs : unique_diff_on 𝕜 s) {x : E} (hx : x ∈ s) {n : ℕ} (hn : (n : ℕ∞) ≤ N) : ‖iterated_fderiv_within 𝕜 n (λ y, f y * g y) s x‖ ≤ ∑ i in finset.range (n+1), (n.choose i : ℝ) * ‖iterated_fderiv_within 𝕜 i f s x‖ * ‖iterated_fderiv_within 𝕜 (n-i) g s x‖ := @@ -2605,8 +2615,8 @@ lemma norm_iterated_fderiv_within_mul_le (continuous_linear_map.op_norm_mul_le _ _) lemma norm_iterated_fderiv_mul_le - {f : E → A} {g : E → A} {N : with_top ℕ} (hf : cont_diff 𝕜 N f) (hg : cont_diff 𝕜 N g) - (x : E) {n : ℕ} (hn : (n : with_top ℕ) ≤ N) : + {f : E → A} {g : E → A} {N : ℕ∞} (hf : cont_diff 𝕜 N f) (hg : cont_diff 𝕜 N g) + (x : E) {n : ℕ} (hn : (n : ℕ∞) ≤ N) : ‖iterated_fderiv 𝕜 n (λ y, f y * g y) x‖ ≤ ∑ i in finset.range (n+1), (n.choose i : ℝ) * ‖iterated_fderiv 𝕜 i f x‖ * ‖iterated_fderiv 𝕜 (n-i) g x‖ := @@ -2618,6 +2628,209 @@ end end +/-- If the derivatives within a set of `g` at `f x` are bounded by `C`, and the `i`-th derivative +within a set of `f` at `x` is bounded by `D^i` for all `1 ≤ i ≤ n`, then the `n`-th derivative +of `g ∘ f` is bounded by `n! * C * D^n`. +This lemma proves this estimate assuming additionally that two of the spaces live in the same +universe, to make an induction possible. Use instead `norm_iterated_fderiv_within_comp_le` that +removes this assumption. -/ +lemma norm_iterated_fderiv_within_comp_le_aux + {Fu Gu : Type u} [normed_add_comm_group Fu] [normed_space 𝕜 Fu] + [normed_add_comm_group Gu] [normed_space 𝕜 Gu] + {g : Fu → Gu} {f : E → Fu} {n : ℕ} {s : set E} {t : set Fu} {x : E} + (hg : cont_diff_on 𝕜 n g t) (hf : cont_diff_on 𝕜 n f s) + (ht : unique_diff_on 𝕜 t) (hs : unique_diff_on 𝕜 s) + (hst : maps_to f s t) (hx : x ∈ s) + {C : ℝ} {D : ℝ} (hC : ∀ i, i ≤ n → ‖iterated_fderiv_within 𝕜 i g t (f x)‖ ≤ C) + (hD : ∀ i, 1 ≤ i → i ≤ n → ‖iterated_fderiv_within 𝕜 i f s x‖ ≤ D^i) : + ‖iterated_fderiv_within 𝕜 n (g ∘ f) s x‖ ≤ n! * C * D^n := +begin + /- We argue by induction on `n`, using that `D^(n+1) (g ∘ f) = D^n (g ' ∘ f ⬝ f')`. The successive + derivatives of `g' ∘ f` are controlled thanks to the inductive assumption, and those of `f'` are + controlled by assumption. + As composition of linear maps is a bilinear map, one may use + `continuous_linear_map.norm_iterated_fderiv_le_of_bilinear_of_le_one` to get from these a bound + on `D^n (g ' ∘ f ⬝ f')`. -/ + unfreezingI { induction n using nat.case_strong_induction_on with n IH generalizing Gu }, + { simpa only [norm_iterated_fderiv_within_zero, nat.factorial_zero, algebra_map.coe_one, + one_mul, pow_zero, mul_one] using hC 0 le_rfl }, + have M : (n : ℕ∞) < n.succ := nat.cast_lt.2 n.lt_succ_self, + have Cnonneg : 0 ≤ C := (norm_nonneg _).trans (hC 0 bot_le), + have Dnonneg : 0 ≤ D, + { have : 1 ≤ n+1, by simp only [le_add_iff_nonneg_left, zero_le'], + simpa only [pow_one] using (norm_nonneg _).trans (hD 1 le_rfl this) }, + -- use the inductive assumption to bound the derivatives of `g' ∘ f`. + have I : ∀ i ∈ finset.range (n+1), + ‖iterated_fderiv_within 𝕜 i ((fderiv_within 𝕜 g t) ∘ f) s x‖ ≤ i! * C * D^i, + { assume i hi, + simp only [finset.mem_range_succ_iff] at hi, + apply IH i hi, + apply hf.of_le (nat.cast_le.2 (hi.trans n.le_succ)), + { assume j hj h'j, + exact hD j hj (h'j.trans (hi.trans n.le_succ)) }, + { apply hg.fderiv_within ht, + simp only [nat.cast_succ], + exact add_le_add_right (nat.cast_le.2 hi) _ }, + { assume j hj, + have : ‖iterated_fderiv_within 𝕜 j (fderiv_within 𝕜 g t) t (f x)‖ + = ‖iterated_fderiv_within 𝕜 (j+1) g t (f x)‖, + by rw [iterated_fderiv_within_succ_eq_comp_right ht (hst hx), linear_isometry_equiv.norm_map], + rw this, + exact hC (j+1) (add_le_add (hj.trans hi) le_rfl) } }, + -- reformulate `hD` as a bound for the derivatives of `f'`. + have J : ∀ i, ‖iterated_fderiv_within 𝕜 (n - i) (fderiv_within 𝕜 f s) s x‖ ≤ D ^ (n - i + 1), + { assume i, + have : ‖iterated_fderiv_within 𝕜 (n - i) (fderiv_within 𝕜 f s) s x‖ + = ‖iterated_fderiv_within 𝕜 (n - i + 1) f s x‖, + by rw [iterated_fderiv_within_succ_eq_comp_right hs hx, linear_isometry_equiv.norm_map], + rw this, + apply hD, + { simp only [le_add_iff_nonneg_left, zero_le'] }, + { apply nat.succ_le_succ tsub_le_self } }, + -- Now put these together: first, notice that we have to bound `D^n (g' ∘ f ⬝ f')`. + calc + ‖iterated_fderiv_within 𝕜 (n+1) (g ∘ f) s x‖ = + ‖iterated_fderiv_within 𝕜 n (λ (y : E), fderiv_within 𝕜 (g ∘ f) s y) s x‖ : + by rw [iterated_fderiv_within_succ_eq_comp_right hs hx, linear_isometry_equiv.norm_map] + ... = ‖iterated_fderiv_within 𝕜 n (λ (y : E), continuous_linear_map.compL 𝕜 E Fu Gu + (fderiv_within 𝕜 g t (f y)) (fderiv_within 𝕜 f s y)) s x‖ : + begin + have L : (1 : ℕ∞) ≤ n.succ, by simpa only [enat.coe_one, nat.one_le_cast] using n.succ_pos, + congr' 1, + apply iterated_fderiv_within_congr hs (λ y hy, _) hx, + apply fderiv_within.comp _ _ _ hst (hs y hy), + { exact hg.differentiable_on L _ (hst hy) }, + { exact hf.differentiable_on L _ hy } + end + -- bound it using the fact that the composition of linear maps is a bilinear operation, + -- for which we have bounds for the`n`-th derivative. + ... ≤ ∑ i in finset.range (n+1), (n.choose i : ℝ) * + ‖iterated_fderiv_within 𝕜 i ((fderiv_within 𝕜 g t) ∘ f) s x‖ + * ‖iterated_fderiv_within 𝕜 (n-i) (fderiv_within 𝕜 f s) s x‖ : + begin + have A : cont_diff_on 𝕜 n ((fderiv_within 𝕜 g t) ∘ f) s, + { apply cont_diff_on.comp _ (hf.of_le M.le) hst, + apply hg.fderiv_within ht, + simp only [nat.cast_succ, le_refl] }, + have B : cont_diff_on 𝕜 n (fderiv_within 𝕜 f s) s, + { apply hf.fderiv_within hs, + simp only [nat.cast_succ, le_refl] }, + exact (continuous_linear_map.compL 𝕜 E Fu Gu) + .norm_iterated_fderiv_within_le_of_bilinear_of_le_one A B hs hx + le_rfl (continuous_linear_map.norm_compL_le 𝕜 E Fu Gu), + end + -- bound each of the terms using the estimates on previous derivatives (that use the inductive + -- assumption for `g' ∘ f`). + ... ≤ ∑ i in finset.range (n+1), (n.choose i : ℝ) * (i! * C * D^i) * (D^(n-i+1)) : + begin + apply finset.sum_le_sum (λ i hi, _), + simp only [mul_assoc (n.choose i : ℝ)], + refine mul_le_mul_of_nonneg_left _ (nat.cast_nonneg _), + apply mul_le_mul (I i hi) (J i) (norm_nonneg _), + positivity, + end + -- We are left with trivial algebraic manipulations to see that this is smaller than + -- the claimed bound. + ... = ∑ i in finset.range (n+1), (n! : ℝ) * (i!⁻¹ * i!) * C * (D^i * D^(n-i+1)) * (n-i)!⁻¹ : + begin + apply finset.sum_congr rfl (λ i hi, _), + simp only [nat.cast_choose ℝ (finset.mem_range_succ_iff.1 hi), div_eq_inv_mul, mul_inv], + ring, + end + ... = ∑ i in finset.range (n+1), (n! : ℝ) * 1 * C * D^(n+1) * (n-i)!⁻¹ : + begin + apply finset.sum_congr rfl (λ i hi, _), + congr' 2, + { congr, + apply inv_mul_cancel, + simpa only [ne.def, nat.cast_eq_zero] using i.factorial_ne_zero }, + { rw ← pow_add, + congr' 1, + rw [nat.add_succ, nat.succ_inj'], + exact nat.add_sub_of_le (finset.mem_range_succ_iff.1 hi) } + end + ... ≤ ∑ i in finset.range (n+1), (n! : ℝ) * 1 * C * D^(n+1) * 1 : + begin + apply finset.sum_le_sum (λ i hi, _), + refine mul_le_mul_of_nonneg_left _ (by positivity), + apply inv_le_one, + simpa only [nat.one_le_cast] using (n-i).factorial_pos, + end + ... = (n+1)! * C * D^(n+1) : + by simp only [mul_assoc, mul_one, finset.sum_const, finset.card_range, nsmul_eq_mul, + nat.factorial_succ, nat.cast_mul], +end + +/-- If the derivatives within a set of `g` at `f x` are bounded by `C`, and the `i`-th derivative +within a set of `f` at `x` is bounded by `D^i` for all `1 ≤ i ≤ n`, then the `n`-th derivative +of `g ∘ f` is bounded by `n! * C * D^n`. -/ +lemma norm_iterated_fderiv_within_comp_le + {g : F → G} {f : E → F} {n : ℕ} {s : set E} {t : set F} {x : E} {N : ℕ∞} + (hg : cont_diff_on 𝕜 N g t) (hf : cont_diff_on 𝕜 N f s) (hn : (n : ℕ∞) ≤ N) + (ht : unique_diff_on 𝕜 t) (hs : unique_diff_on 𝕜 s) + (hst : maps_to f s t) (hx : x ∈ s) + {C : ℝ} {D : ℝ} (hC : ∀ i, i ≤ n → ‖iterated_fderiv_within 𝕜 i g t (f x)‖ ≤ C) + (hD : ∀ i, 1 ≤ i → i ≤ n → ‖iterated_fderiv_within 𝕜 i f s x‖ ≤ D^i) : + ‖iterated_fderiv_within 𝕜 n (g ∘ f) s x‖ ≤ n! * C * D^n := +begin + /- We reduce the bound to the case where all spaces live in the same universe (in which we + already have proved the result), by using linear isometries between the spaces and their `ulift` + to a common universe. These linear isometries preserve the norm of the iterated derivative. -/ + let Fu : Type (max uF uG) := ulift.{uG uF} F, + let Gu : Type (max uF uG) := ulift.{uF uG} G, + have isoF : Fu ≃ₗᵢ[𝕜] F := linear_isometry_equiv.ulift 𝕜 F, + have isoG : Gu ≃ₗᵢ[𝕜] G := linear_isometry_equiv.ulift 𝕜 G, + -- lift `f` and `g` to versions `fu` and `gu` on the lifted spaces. + let fu : E → Fu := isoF.symm ∘ f, + let gu : Fu → Gu := isoG.symm ∘ g ∘ isoF, + let tu := isoF ⁻¹' t, + have htu : unique_diff_on 𝕜 tu, + from isoF.to_continuous_linear_equiv.unique_diff_on_preimage_iff.2 ht, + have hstu : maps_to fu s tu, + { assume y hy, + simpa only [mem_preimage, linear_isometry_equiv.apply_symm_apply] using hst hy }, + have Ffu : isoF (fu x) = f x, by simp only [linear_isometry_equiv.apply_symm_apply], + -- All norms are preserved by the lifting process. + have hfu : cont_diff_on 𝕜 n fu s, from isoF.symm.cont_diff.comp_cont_diff_on (hf.of_le hn), + have hgu : cont_diff_on 𝕜 n gu tu, from isoG.symm.cont_diff.comp_cont_diff_on + ((hg.of_le hn).comp_continuous_linear_map (isoF : Fu →L[𝕜] F)), + have Nfu : ∀ i, ‖iterated_fderiv_within 𝕜 i fu s x‖ = ‖iterated_fderiv_within 𝕜 i f s x‖, + { assume i, + rw linear_isometry_equiv.norm_iterated_fderiv_within_comp_left _ _ hs hx }, + simp_rw [← Nfu] at hD, + have Ngu : ∀ i, ‖iterated_fderiv_within 𝕜 i gu tu (fu x)‖ + = ‖iterated_fderiv_within 𝕜 i g t (f x)‖, + { assume i, + rw linear_isometry_equiv.norm_iterated_fderiv_within_comp_left _ _ htu (hstu hx), + rw [linear_isometry_equiv.norm_iterated_fderiv_within_comp_right _ _ ht, Ffu], + rw Ffu, + exact hst hx }, + simp_rw [← Ngu] at hC, + have Nfgu : ‖iterated_fderiv_within 𝕜 n (g ∘ f) s x‖ = ‖iterated_fderiv_within 𝕜 n (gu ∘ fu) s x‖, + { have : gu ∘ fu = isoG.symm ∘ g ∘ f, + { ext x, + simp only [comp_app, linear_isometry_equiv.map_eq_iff, + linear_isometry_equiv.apply_symm_apply] }, + rw [this, linear_isometry_equiv.norm_iterated_fderiv_within_comp_left _ _ hs hx] }, + -- deduce the required bound from the one for `gu ∘ fu`. + rw Nfgu, + exact norm_iterated_fderiv_within_comp_le_aux hgu hfu htu hs hstu hx hC hD, +end + +/-- If the derivatives of `g` at `f x` are bounded by `C`, and the `i`-th derivative +of `f` at `x` is bounded by `D^i` for all `1 ≤ i ≤ n`, then the `n`-th derivative +of `g ∘ f` is bounded by `n! * C * D^n`. -/ +lemma norm_iterated_fderiv_comp_le + {g : F → G} {f : E → F} {n : ℕ} {N : ℕ∞} + (hg : cont_diff 𝕜 N g) (hf : cont_diff 𝕜 N f) (hn : (n : ℕ∞) ≤ N) (x : E) + {C : ℝ} {D : ℝ} (hC : ∀ i, i ≤ n → ‖iterated_fderiv 𝕜 i g (f x)‖ ≤ C) + (hD : ∀ i, 1 ≤ i → i ≤ n → ‖iterated_fderiv 𝕜 i f x‖ ≤ D^i) : + ‖iterated_fderiv 𝕜 n (g ∘ f) x‖ ≤ n! * C * D^n := +begin + simp_rw [← iterated_fderiv_within_univ] at ⊢ hC hD, + exact norm_iterated_fderiv_within_comp_le hg.cont_diff_on hf.cont_diff_on hn unique_diff_on_univ + unique_diff_on_univ (maps_to_univ _ _) (mem_univ x) hC hD, +end section apply lemma norm_iterated_fderiv_within_clm_apply {f : E → (F →L[𝕜] G)} {g : E → F} {s : set E} {x : E} From 98bd247d933fb581ff37244a5998bd33d81dd46d Mon Sep 17 00:00:00 2001 From: Aaron Anderson Date: Tue, 25 Apr 2023 12:54:02 +0000 Subject: [PATCH 0871/1291] feat(model_theory/types): Realized types (#17848) Defines `first_order.language.Theory.type_of`, the type of a given tuple. Defines what it means for a type to be realized. Co-authored-by: Aaron Anderson <65780815+awainverse@users.noreply.github.com> --- src/model_theory/types.lean | 53 +++++++++++++++++++++++++++++++++++++ 1 file changed, 53 insertions(+) diff --git a/src/model_theory/types.lean b/src/model_theory/types.lean index ee7fe94891468..2df5368d2a6f2 100644 --- a/src/model_theory/types.lean +++ b/src/model_theory/types.lean @@ -13,10 +13,15 @@ This file defines the space of complete types over a first-order theory. ## Main Definitions * `first_order.language.Theory.complete_type`: `T.complete_type α` consists of complete types over the theory `T` with variables `α`. +* `first_order.language.Theory.type_of` is the type of a given tuple. +* `first_order.language.Theory.realized_types`: `T.realized_types M α` is the set of + types in `T.complete_type α` that are realized in `M` - that is, the type of some tuple in `M`. ## Main Results * `first_order.language.Theory.complete_type.nonempty_iff`: The space `T.complete_type α` is nonempty exactly when `T` is satisfiable. +* `first_order.language.Theory.complete_type.exists_Model_is_realized_in`: Every type is realized in +some model. ## Implementation Notes * Complete types are implemented as maximal consistent theories in an expanded language. @@ -151,6 +156,54 @@ end end complete_type +variables {M : Type w'} [L.Structure M] [nonempty M] [M ⊨ T] (T) + +/-- The set of all formulas true at a tuple in a structure forms a complete type. -/ +def type_of (v : α → M) : T.complete_type α := +begin + haveI : (constants_on α).Structure M := constants_on.Structure v, + exact { to_Theory := L[[α]].complete_theory M, + subset' := model_iff_subset_complete_theory.1 ((Lhom.on_Theory_model _ T).2 infer_instance), + is_maximal' := complete_theory.is_maximal _ _ }, +end + +namespace complete_type + +variables {T} {v : α → M} + +@[simp] lemma mem_type_of {φ : L[[α]].sentence} : + φ ∈ T.type_of v ↔ (formula.equiv_sentence.symm φ).realize v := +begin + letI : (constants_on α).Structure M := constants_on.Structure v, + exact mem_complete_theory.trans (formula.realize_equiv_sentence_symm _ _ _).symm, +end + +lemma formula_mem_type_of {φ : L.formula α} : + formula.equiv_sentence φ ∈ T.type_of v ↔ φ.realize v := +by simp + +end complete_type + +variable (M) + +/-- A complete type `p` is realized in a particular structure when there is some + tuple `v` whose type is `p`. -/ +@[simp] def realized_types (α : Type w) : set (T.complete_type α) := +set.range (T.type_of : (α → M) → T.complete_type α) + +theorem exists_Model_is_realized_in (p : T.complete_type α) : + ∃ (M : Theory.Model.{u v (max u v w)} T), + p ∈ T.realized_types M α := +begin + obtain ⟨M⟩ := p.is_maximal.1, + refine ⟨(M.subtheory_Model p.subset).reduct (L.Lhom_with_constants α), (λ a, (L.con a : M)), _⟩, + refine set_like.ext (λ φ, _), + simp only [complete_type.mem_type_of], + refine (formula.realize_equiv_sentence_symm_con _ _).trans (trans (trans _ + (p.is_maximal.is_complete.realize_sentence_iff φ M)) (p.is_maximal.mem_iff_models φ).symm), + refl, +end + end Theory end language end first_order From d2d964c64f8ddcccd6704a731c41f95d13e72f5c Mon Sep 17 00:00:00 2001 From: Floris van Doorn Date: Tue, 25 Apr 2023 17:43:10 +0000 Subject: [PATCH 0872/1291] feat(topology/vector_bundle/basic): define in_coordinates (#18826) * From the sphere eversion project --- src/topology/vector_bundle/basic.lean | 77 +++++++++++++++++++++++++++ src/topology/vector_bundle/hom.lean | 31 +++++++++-- 2 files changed, 103 insertions(+), 5 deletions(-) diff --git a/src/topology/vector_bundle/basic.lean b/src/topology/vector_bundle/basic.lean index 364f13df75491..67da394c060e5 100644 --- a/src/topology/vector_bundle/basic.lean +++ b/src/topology/vector_bundle/basic.lean @@ -29,6 +29,21 @@ If these conditions are satisfied, we register the typeclass `vector_bundle R F We define constructions on vector bundles like pullbacks and direct sums in other files. +## Main Definitions + +* `trivialization.is_linear`: a class stating that a trivialization is fiberwise linear on its base + set. +* `trivialization.linear_equiv_at` and `trivialization.continuous_linear_map_at` are the + (continuous) linear fiberwise equivalences a trivialization induces. +* They have forward maps `trivialization.linear_map_at` / `trivialization.continuous_linear_map_at` + and inverses `trivialization.symmₗ` / `trivialization.symmL`. Note that these are all defined + everywhere, since they are extended using the zero function. +* `trivialization.coord_changeL` is the coordinate change induced by two trivializations. It only + makes sense on the intersection of their base sets, but is extended outside it using the identity. +* Given a continuous (semi)linear map between `E x` and `E' y` where `E` and `E'` are bundles over + possibly different base sets, `continuous_linear_map.in_coordinates` turns this into a continuous + (semi)linear map between the chosen fibers of those bundles. + ## Implementation notes The implementation choices in the vector bundle definition are discussed in the "Implementation @@ -868,4 +883,66 @@ lemma to_vector_bundle : end vector_prebundle +namespace continuous_linear_map +variables {𝕜₁ 𝕜₂ : Type*} [nontrivially_normed_field 𝕜₁] [nontrivially_normed_field 𝕜₂] +variables {σ : 𝕜₁ →+* 𝕜₂} +variables {B' : Type*} [topological_space B'] + +variables [normed_space 𝕜₁ F] [Π x, module 𝕜₁ (E x)] [topological_space (total_space E)] +variables {F' : Type*} [normed_add_comm_group F'] [normed_space 𝕜₂ F'] + {E' : B' → Type*} [Π x, add_comm_monoid (E' x)] [Π x, module 𝕜₂ (E' x)] + [topological_space (total_space E')] +variables [Π x, topological_space (E x)] [fiber_bundle F E] [vector_bundle 𝕜₁ F E] +variables [Π x, topological_space (E' x)] [fiber_bundle F' E'] [vector_bundle 𝕜₂ F' E'] +variables (F E F' E') + +/-- When `ϕ` is a continuous (semi)linear map between the fibers `E x` and `E' y` of two vector +bundles `E` and `E'`, `continuous_linear_map.in_coordinates F E F' E' x₀ x y₀ y ϕ` is a coordinate +change of this continuous linear map w.r.t. the chart around `x₀` and the chart around `y₀`. + +It is defined by composing `ϕ` with appropriate coordinate changes given by the vector bundles +`E` and `E'`. +We use the operations `trivialization.continuous_linear_map_at` and `trivialization.symmL` in the +definition, instead of `trivialization.continuous_linear_equiv_at`, so that +`continuous_linear_map.in_coordinates` is defined everywhere (but see +`continuous_linear_map.in_coordinates_eq`). + +This is the (second component of the) underlying function of a trivialization of the hom-bundle +(see `hom_trivialization_at_apply`). However, note that `continuous_linear_map.in_coordinates` is +defined even when `x` and `y` live in different base sets. +Therefore, it is is also convenient when working with the hom-bundle between pulled back bundles. +-/ +def in_coordinates (x₀ x : B) (y₀ y : B') (ϕ : E x →SL[σ] E' y) : F →SL[σ] F' := +((trivialization_at F' E' y₀).continuous_linear_map_at 𝕜₂ y).comp $ ϕ.comp $ +(trivialization_at F E x₀).symmL 𝕜₁ x + +variables {F F'} + +/-- rewrite `in_coordinates` using continuous linear equivalences. -/ +lemma in_coordinates_eq (x₀ x : B) (y₀ y : B') (ϕ : E x →SL[σ] E' y) + (hx : x ∈ (trivialization_at F E x₀).base_set) + (hy : y ∈ (trivialization_at F' E' y₀).base_set) : + in_coordinates F E F' E' x₀ x y₀ y ϕ = + ((trivialization_at F' E' y₀).continuous_linear_equiv_at 𝕜₂ y hy : E' y →L[𝕜₂] F').comp (ϕ.comp $ + (((trivialization_at F E x₀).continuous_linear_equiv_at 𝕜₁ x hx).symm : F →L[𝕜₁] E x)) := +begin + ext, + simp_rw [in_coordinates, continuous_linear_map.coe_comp', continuous_linear_equiv.coe_coe, + trivialization.coe_continuous_linear_equiv_at_eq, + trivialization.symm_continuous_linear_equiv_at_eq] +end + +/-- rewrite `in_coordinates` in a `vector_bundle_core`. -/ +protected lemma vector_bundle_core.in_coordinates_eq {ι ι'} (Z : vector_bundle_core 𝕜₁ B F ι) + (Z' : vector_bundle_core 𝕜₂ B' F' ι') + {x₀ x : B} {y₀ y : B'} (ϕ : F →SL[σ] F') + (hx : x ∈ Z.base_set (Z.index_at x₀)) + (hy : y ∈ Z'.base_set (Z'.index_at y₀)) : + in_coordinates F Z.fiber F' Z'.fiber x₀ x y₀ y ϕ = + (Z'.coord_change (Z'.index_at y) (Z'.index_at y₀) y).comp (ϕ.comp $ + Z.coord_change (Z.index_at x₀) (Z.index_at x) x) := +by simp_rw [in_coordinates, Z'.trivialization_at_continuous_linear_map_at hy, + Z.trivialization_at_symmL hx] + +end continuous_linear_map end diff --git a/src/topology/vector_bundle/hom.lean b/src/topology/vector_bundle/hom.lean index e37608d990d0a..20c00d4d11894 100644 --- a/src/topology/vector_bundle/hom.lean +++ b/src/topology/vector_bundle/hom.lean @@ -45,9 +45,9 @@ variables {𝕜₁ 𝕜₂ : Type*} [normed_field 𝕜₁] [normed_field 𝕜₂ variables (σ : 𝕜₁ →+* 𝕜₂) variables {B : Type*} variables (F₁ : Type*) (E₁ : B → Type*) [Π x, add_comm_monoid (E₁ x)] [Π x, module 𝕜₁ (E₁ x)] -variables [Π x : B, topological_space (E₁ x)] +variables [Π x, topological_space (E₁ x)] variables (F₂ : Type*) (E₂ : B → Type*) [Π x, add_comm_monoid (E₂ x)] [Π x, module 𝕜₂ (E₂ x)] -variables [Π x : B, topological_space (E₂ x)] +variables [Π x, topological_space (E₂ x)] include F₁ F₂ @@ -61,7 +61,7 @@ include F₁ F₂ We intentionally add `F₁` and `F₂` as arguments to this type, so that instances on this type (that depend on `F₁` and `F₂`) actually refer to `F₁` and `F₂`. -/ @[derive inhabited, nolint unused_arguments] -def bundle.continuous_linear_map (x : B) : Type* := +protected def bundle.continuous_linear_map (x : B) : Type* := E₁ x →SL[σ] E₂ x instance bundle.continuous_linear_map.add_monoid_hom_class (x : B) : @@ -110,8 +110,8 @@ def continuous_linear_map_coord_change (F₁ →SL[σ] F₂) ≃L[𝕜₂] F₁ →SL[σ] F₂) variables {σ e₁ e₁' e₂ e₂'} -variables [Π x : B, topological_space (E₁ x)] [fiber_bundle F₁ E₁] -variables [Π x : B, topological_space (E₂ x)] [fiber_bundle F₂ E₂] +variables [Π x, topological_space (E₁ x)] [fiber_bundle F₁ E₁] +variables [Π x, topological_space (E₂ x)] [fiber_bundle F₂ E₂] lemma continuous_on_continuous_linear_map_coord_change [vector_bundle 𝕜₁ F₁ E₁] [vector_bundle 𝕜₂ F₂ E₂] @@ -315,3 +315,24 @@ lemma trivialization.continuous_linear_map_apply e₁.continuous_linear_map σ e₂ p = ⟨p.1, (e₂.continuous_linear_map_at 𝕜₂ p.1).comp $ p.2.comp $ e₁.symmL 𝕜₁ p.1⟩ := rfl + +omit he₁ he₂ + +lemma hom_trivialization_at_apply (x₀ : B) + (x : total_space (bundle.continuous_linear_map σ F₁ E₁ F₂ E₂)) : + trivialization_at (F₁ →SL[σ] F₂) (bundle.continuous_linear_map σ F₁ E₁ F₂ E₂) x₀ x = + ⟨x.1, in_coordinates F₁ E₁ F₂ E₂ x₀ x.1 x₀ x.1 x.2⟩ := +rfl + +@[simp, mfld_simps] +lemma hom_trivialization_at_source (x₀ : B) : + (trivialization_at (F₁ →SL[σ] F₂) (bundle.continuous_linear_map σ F₁ E₁ F₂ E₂) x₀).source = + π (bundle.continuous_linear_map σ F₁ E₁ F₂ E₂) ⁻¹' + ((trivialization_at F₁ E₁ x₀).base_set ∩ (trivialization_at F₂ E₂ x₀).base_set) := +rfl + +@[simp, mfld_simps] +lemma hom_trivialization_at_target (x₀ : B) : + (trivialization_at (F₁ →SL[σ] F₂) (bundle.continuous_linear_map σ F₁ E₁ F₂ E₂) x₀).target = + ((trivialization_at F₁ E₁ x₀).base_set ∩ (trivialization_at F₂ E₂ x₀).base_set) ×ˢ set.univ := +rfl From efe03a53241aaa777c1016a7a0e71dd3b92a4313 Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Tue, 25 Apr 2023 21:34:29 +0000 Subject: [PATCH 0873/1291] fix(topology/continuous_function/algebra): let `continuous_map.module'` use a `semiring` instead of ` ring` (#18868) This fixes an issue where Lean couldn't recognize via TC inference that, for a topological semiring `R`, that `C(R, R)` is an `R`-module, despite the fact that it knows it is an `R`-algebra through `continuous_map.algebra`. All that is necessary is to weaken the type class assumptions. --- src/topology/continuous_function/algebra.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/topology/continuous_function/algebra.lean b/src/topology/continuous_function/algebra.lean index bf4e1f25a46af..24f792af517d2 100644 --- a/src/topology/continuous_function/algebra.lean +++ b/src/topology/continuous_function/algebra.lean @@ -790,7 +790,7 @@ instance has_smul' {α : Type*} [topological_space α] ⟨λ f g, ⟨λ x, (f x) • (g x), (continuous.smul f.2 g.2)⟩⟩ instance module' {α : Type*} [topological_space α] - (R : Type*) [ring R] [topological_space R] [topological_ring R] + (R : Type*) [semiring R] [topological_space R] [topological_semiring R] (M : Type*) [topological_space M] [add_comm_monoid M] [has_continuous_add M] [module R M] [has_continuous_smul R M] : module C(α, R) C(α, M) := From 52fa514ec337dd970d71d8de8d0fd68b455a1e54 Mon Sep 17 00:00:00 2001 From: Sam Ezeh Date: Tue, 25 Apr 2023 23:03:08 +0000 Subject: [PATCH 0874/1291] feat(data/set/finite): Add lemmas relating definitions of infinite sets (#18620) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Prove the following lemmas (and their dual) * `set.infinite_of_forall_exists_lt`: `(∀ a, ∃ b ∈ s, a < b) → s.infinite` in a nonempty preorder * `set.infinite_iff_exists_lt`: `(∀ a, ∃ b ∈ s, a < b) ↔ s.infinite` in a locally finite order with a bottom element Co-authored-by: Yaël Dillies Co-authored-by: Eric Wieser --- src/data/finset/locally_finite.lean | 28 ++++++++++++++++++++++++++++ src/data/nat/lattice.lean | 3 ++- src/data/nat/nth.lean | 2 +- src/data/set/finite.lean | 20 +++++++++++++++++--- src/group_theory/exponent.lean | 2 +- 5 files changed, 49 insertions(+), 6 deletions(-) diff --git a/src/data/finset/locally_finite.lean b/src/data/finset/locally_finite.lean index 7366deb5165fa..8dc178e34bc1d 100644 --- a/src/data/finset/locally_finite.lean +++ b/src/data/finset/locally_finite.lean @@ -259,6 +259,9 @@ lemma Ioi_subset_Ici_self : Ioi a ⊆ Ici a := by simpa [←coe_subset] using se lemma _root_.bdd_below.finite {s : set α} (hs : bdd_below s) : s.finite := let ⟨a, ha⟩ := hs in (Ici a).finite_to_set.subset $ λ x hx, mem_Ici.2 $ ha hx +lemma _root_.set.infinite.not_bdd_below {s : set α} : s.infinite → ¬ bdd_below s := +mt bdd_below.finite + variables [fintype α] lemma filter_lt_eq_Ioi [decidable_pred ((<) a)] : univ.filter ((<) a) = Ioi a := by { ext, simp } @@ -273,6 +276,9 @@ lemma Iio_subset_Iic_self : Iio a ⊆ Iic a := by simpa [←coe_subset] using se lemma _root_.bdd_above.finite {s : set α} (hs : bdd_above s) : s.finite := hs.dual.finite +lemma _root_.set.infinite.not_bdd_above {s : set α} : s.infinite → ¬ bdd_above s := +mt bdd_above.finite + variables [fintype α] lemma filter_gt_eq_Iio [decidable_pred (< a)] : univ.filter (< a) = Iio a := by { ext, simp } @@ -504,6 +510,28 @@ end end locally_finite_order +section locally_finite_order_bot +variables [locally_finite_order_bot α] {s : set α} + +lemma _root_.set.infinite.exists_gt (hs : s.infinite) : ∀ a, ∃ b ∈ s, a < b := +not_bdd_above_iff.1 hs.not_bdd_above + +lemma _root_.set.infinite_iff_exists_gt [nonempty α] : s.infinite ↔ ∀ a, ∃ b ∈ s, a < b := +⟨set.infinite.exists_gt, set.infinite_of_forall_exists_gt⟩ + +end locally_finite_order_bot + +section locally_finite_order_top +variables [locally_finite_order_top α] {s : set α} + +lemma _root_.set.infinite.exists_lt (hs : s.infinite) : ∀ a, ∃ b ∈ s, b < a := +not_bdd_below_iff.1 hs.not_bdd_below + +lemma _root_.set.infinite_iff_exists_lt [nonempty α] : s.infinite ↔ ∀ a, ∃ b ∈ s, b < a := +⟨set.infinite.exists_lt, set.infinite_of_forall_exists_lt⟩ + +end locally_finite_order_top + variables [fintype α] [locally_finite_order_top α] [locally_finite_order_bot α] lemma Ioi_disj_union_Iio (a : α) : diff --git a/src/data/nat/lattice.lean b/src/data/nat/lattice.lean index 1647489078bf4..61ff4e7e97e7e 100644 --- a/src/data/nat/lattice.lean +++ b/src/data/nat/lattice.lean @@ -3,6 +3,7 @@ Copyright (c) 2018 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Floris van Doorn, Gabriel Ebner, Yury Kudryashov -/ +import data.nat.interval import order.conditionally_complete_lattice.finset /-! @@ -37,7 +38,7 @@ lemma Sup_def {s : set ℕ} (h : ∃n, ∀a∈s, a ≤ n) : dif_pos _ lemma _root_.set.infinite.nat.Sup_eq_zero {s : set ℕ} (h : s.infinite) : Sup s = 0 := -dif_neg $ λ ⟨n, hn⟩, let ⟨k, hks, hk⟩ := h.exists_nat_lt n in (hn k hks).not_lt hk +dif_neg $ λ ⟨n, hn⟩, let ⟨k, hks, hk⟩ := h.exists_gt n in (hn k hks).not_lt hk @[simp] lemma Inf_eq_zero {s : set ℕ} : Inf s = 0 ↔ 0 ∈ s ∨ s = ∅ := begin diff --git a/src/data/nat/nth.lean b/src/data/nat/nth.lean index 9f4da094252c4..18e2db160b439 100644 --- a/src/data/nat/nth.lean +++ b/src/data/nat/nth.lean @@ -294,7 +294,7 @@ begin suffices h : Inf {i : ℕ | p i ∧ n ≤ i} ∈ {i : ℕ | p i ∧ n ≤ i}, { exact h.2 }, apply Inf_mem, - obtain ⟨m, hp, hn⟩ := hp.exists_nat_lt n, + obtain ⟨m, hp, hn⟩ := hp.exists_gt n, exact ⟨m, hp, hn.le⟩ end diff --git a/src/data/set/finite.lean b/src/data/set/finite.lean index 0c891a6ca3202..b7eb2f977fc42 100644 --- a/src/data/set/finite.lean +++ b/src/data/set/finite.lean @@ -1054,9 +1054,6 @@ theorem infinite_of_injective_forall_mem [infinite α] {s : set β} {f : α → (hi : injective f) (hf : ∀ x : α, f x ∈ s) : s.infinite := by { rw ←range_subset_iff at hf, exact (infinite_range_of_injective hi).mono hf } -lemma infinite.exists_nat_lt {s : set ℕ} (hs : s.infinite) (n : ℕ) : ∃ m ∈ s, n < m := -let ⟨m, hm⟩ := (hs.diff $ set.finite_le_nat n).nonempty in ⟨m, by simpa using hm⟩ - lemma infinite.exists_not_mem_finset {s : set α} (hs : s.infinite) (f : finset α) : ∃ a ∈ s, a ∉ f := let ⟨a, has, haf⟩ := (hs.diff (to_finite f)).nonempty @@ -1076,6 +1073,23 @@ end /-! ### Order properties -/ +section preorder +variables [preorder α] [nonempty α] {s : set α} + +lemma infinite_of_forall_exists_gt (h : ∀ a, ∃ b ∈ s, a < b) : s.infinite := +begin + inhabit α, + set f : ℕ → α := λ n, nat.rec_on n (h default).some (λ n a, (h a).some), + have hf : ∀ n, f n ∈ s := by rintro (_ | _); exact (h _).some_spec.some, + refine infinite_of_injective_forall_mem (strict_mono_nat_of_lt_succ $ λ n, _).injective hf, + exact (h _).some_spec.some_spec, +end + +lemma infinite_of_forall_exists_lt (h : ∀ a, ∃ b ∈ s, b < a) : s.infinite := +@infinite_of_forall_exists_gt αᵒᵈ _ _ _ h + +end preorder + lemma finite_is_top (α : Type*) [partial_order α] : {x : α | is_top x}.finite := (subsingleton_is_top α).finite diff --git a/src/group_theory/exponent.lean b/src/group_theory/exponent.lean index 153276863eb4a..f77a742a26313 100644 --- a/src/group_theory/exponent.lean +++ b/src/group_theory/exponent.lean @@ -196,7 +196,7 @@ variable {G} begin refine ⟨λ he, _, λ he, _⟩, { by_contra h, - obtain ⟨m, ⟨t, rfl⟩, het⟩ := set.infinite.exists_nat_lt h (exponent G), + obtain ⟨m, ⟨t, rfl⟩, het⟩ := set.infinite.exists_gt h (exponent G), exact pow_ne_one_of_lt_order_of' he het (pow_exponent_eq_one t) }, { lift (set.range order_of) to finset ℕ using he with t ht, have htpos : 0 < t.prod id, From f0b3759a8ef0bd8239ecdaa5e1089add5feebe1a Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Tue, 25 Apr 2023 23:03:09 +0000 Subject: [PATCH 0875/1291] chore(set_theory/zfc/basic): reverse `sInter_coe` (#18773) It now matches `coe_sUnion`, and works as a `norm_cast` lemma. Mathlib 4: https://github.com/leanprover-community/mathlib4/pull/3345 --- src/set_theory/zfc/basic.lean | 5 +++-- 1 file changed, 3 insertions(+), 2 deletions(-) diff --git a/src/set_theory/zfc/basic.lean b/src/set_theory/zfc/basic.lean index 062030c1039ae..78ba1fc1517e3 100644 --- a/src/set_theory/zfc/basic.lean +++ b/src/set_theory/zfc/basic.lean @@ -1114,8 +1114,9 @@ begin exact H _ hxz end -@[simp, norm_cast] theorem sInter_coe {x : Set.{u}} (h : x.nonempty) : ⋂₀ (x : Class.{u}) = ⋂₀ x := -set.ext $ λ y, sInter_apply.trans (Set.mem_sInter h).symm +@[simp, norm_cast] theorem coe_sInter {x : Set.{u}} (h : x.nonempty) : + ↑(⋂₀ x) = ⋂₀ (x : Class.{u}) := +set.ext $ λ y, (Set.mem_sInter h).trans sInter_apply.symm theorem mem_of_mem_sInter {x y z : Class} (hy : y ∈ ⋂₀ x) (hz : z ∈ x) : y ∈ z := by { obtain ⟨w, rfl, hw⟩ := hy, exact coe_mem.2 (hw z hz) } From 8000bbbe2e9d39b84edb993d88781f536a8a3fa8 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 26 Apr 2023 01:46:20 +0000 Subject: [PATCH 0876/1291] feat(analysis/normed_space/basic): spheres have no interior (#18869) This follows the pattern set by the nearby lemmas of having a primed and unprimed version. --- src/analysis/normed_space/basic.lean | 16 ++++++++++++++++ src/topology/metric_space/basic.lean | 3 +++ 2 files changed, 19 insertions(+) diff --git a/src/analysis/normed_space/basic.lean b/src/analysis/normed_space/basic.lean index 1b623e6761747..830b4367d1f56 100644 --- a/src/analysis/normed_space/basic.lean +++ b/src/analysis/normed_space/basic.lean @@ -187,6 +187,14 @@ theorem frontier_closed_ball [normed_space ℝ E] (x : E) {r : ℝ} (hr : r ≠ by rw [frontier, closure_closed_ball, interior_closed_ball x hr, closed_ball_diff_ball] +theorem interior_sphere [normed_space ℝ E] (x : E) {r : ℝ} (hr : r ≠ 0) : + interior (sphere x r) = ∅ := +by rw [←frontier_closed_ball x hr, interior_frontier is_closed_ball] + +theorem frontier_sphere [normed_space ℝ E] (x : E) {r : ℝ} (hr : r ≠ 0) : + frontier (sphere x r) = sphere x r := +by rw [is_closed_sphere.frontier_eq, interior_sphere x hr, diff_empty] + instance {E : Type*} [normed_add_comm_group E] [normed_space ℚ E] (e : E) : discrete_topology $ add_subgroup.zmultiples e := begin @@ -396,6 +404,14 @@ theorem frontier_closed_ball' [normed_space ℝ E] [nontrivial E] (x : E) (r : frontier (closed_ball x r) = sphere x r := by rw [frontier, closure_closed_ball, interior_closed_ball' x r, closed_ball_diff_ball] +@[simp] theorem interior_sphere' [normed_space ℝ E] [nontrivial E] (x : E) (r : ℝ) : + interior (sphere x r) = ∅ := +by rw [←frontier_closed_ball' x, interior_frontier is_closed_ball] + +@[simp] theorem frontier_sphere' [normed_space ℝ E] [nontrivial E] (x : E) (r : ℝ) : + frontier (sphere x r) = sphere x r := +by rw [is_closed_sphere.frontier_eq, interior_sphere' x, diff_empty] + variables {α} lemma rescale_to_shell_zpow {c : α} (hc : 1 < ‖c‖) {ε : ℝ} (εpos : 0 < ε) {x : E} (hx : x ≠ 0) : diff --git a/src/topology/metric_space/basic.lean b/src/topology/metric_space/basic.lean index ab79a92ec57ad..4a34ea24186fa 100644 --- a/src/topology/metric_space/basic.lean +++ b/src/topology/metric_space/basic.lean @@ -1647,6 +1647,9 @@ is_closed_eq (continuous_id.dist continuous_const) continuous_const @[simp] theorem closure_closed_ball : closure (closed_ball x ε) = closed_ball x ε := is_closed_ball.closure_eq +@[simp] theorem closure_sphere : closure (sphere x ε) = sphere x ε := +is_closed_sphere.closure_eq + theorem closure_ball_subset_closed_ball : closure (ball x ε) ⊆ closed_ball x ε := closure_minimal ball_subset_closed_ball is_closed_ball From 83df6d6ebd4a43b472501c515516a37a9e3d7503 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 26 Apr 2023 08:09:41 +0000 Subject: [PATCH 0877/1291] feat(measure_theory/measure/haar_of_basis): put the canonical measure on euclidean space (#18870) --- src/measure_theory/measure/haar_of_basis.lean | 23 +++++++++++++++++++ 1 file changed, 23 insertions(+) diff --git a/src/measure_theory/measure/haar_of_basis.lean b/src/measure_theory/measure/haar_of_basis.lean index ae4878316a13c..cbb1fe7460ded 100644 --- a/src/measure_theory/measure/haar_of_basis.lean +++ b/src/measure_theory/measure/haar_of_basis.lean @@ -5,6 +5,7 @@ Authors: Sébastien Gouëzel -/ import measure_theory.measure.haar import analysis.inner_product_space.pi_L2 +import measure_theory.constructions.pi /-! # Additive Haar measure constructed from a basis @@ -159,3 +160,25 @@ is proved in `orthonormal_basis.volume_parallelepiped`. -/ /- This instance should not be necessary, but Lean has difficulties to find it in product situations if we do not declare it explicitly. -/ instance real.measure_space : measure_space ℝ := by apply_instance + +/-! # Miscellaneous instances for `euclidean_space` + +In combination with `measure_space_of_inner_product_space`, these put a `measure_space` structure +on `euclidean_space`. -/ +namespace euclidean_space +variables (ι) + +-- TODO: do we want these instances for `pi_Lp` too? +instance : measurable_space (euclidean_space ℝ ι) := measurable_space.pi +instance : borel_space (euclidean_space ℝ ι) := pi.borel_space + +/-- `pi_Lp.equiv` as a `measurable_equiv`. -/ +@[simps to_equiv] +protected def measurable_equiv : euclidean_space ℝ ι ≃ᵐ (ι → ℝ) := +{ to_equiv := pi_Lp.equiv _ _, + measurable_to_fun := measurable_id, + measurable_inv_fun := measurable_id } + +lemma coe_measurable_equiv : ⇑(euclidean_space.measurable_equiv ι) = pi_Lp.equiv 2 _ := rfl + +end euclidean_space From b5665fd3fb2a80ee05ff42b6031ef2055b8f9d85 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 26 Apr 2023 09:56:01 +0000 Subject: [PATCH 0878/1291] feat(data/matrix/block): injectivity lemmas (#18842) Block matrices are equal if their blocks are equal --- src/data/matrix/block.lean | 29 +++++++++++++++++++++++++++++ 1 file changed, 29 insertions(+) diff --git a/src/data/matrix/block.lean b/src/data/matrix/block.lean index bc311dd65413f..0ae6e9748f3e0 100644 --- a/src/data/matrix/block.lean +++ b/src/data/matrix/block.lean @@ -111,6 +111,19 @@ rfl (from_blocks A B C D).to_blocks₂₂ = D := rfl +/-- Two block matrices are equal if their blocks are equal. -/ +lemma ext_iff_blocks {A B : matrix (n ⊕ o) (l ⊕ m) α} : + A = B ↔ A.to_blocks₁₁ = B.to_blocks₁₁ ∧ A.to_blocks₁₂ = B.to_blocks₁₂ ∧ + A.to_blocks₂₁ = B.to_blocks₂₁ ∧ A.to_blocks₂₂ = B.to_blocks₂₂ := +⟨λ h, h ▸ ⟨rfl, rfl, rfl, rfl⟩, λ ⟨h₁₁, h₁₂, h₂₁, h₂₂⟩, + by rw [←from_blocks_to_blocks A, ←from_blocks_to_blocks B, h₁₁, h₁₂, h₂₁, h₂₂]⟩ + +@[simp] lemma from_blocks_inj + {A : matrix n l α} {B : matrix n m α} {C : matrix o l α} {D : matrix o m α} + {A' : matrix n l α} {B' : matrix n m α} {C' : matrix o l α} {D' : matrix o m α} : + from_blocks A B C D = from_blocks A' B' C' D' ↔ A = A' ∧ B = B' ∧ C = C' ∧ D = D' := +ext_iff_blocks + lemma from_blocks_map (A : matrix n l α) (B : matrix n m α) (C : matrix o l α) (D : matrix o m α) (f : α → β) : (from_blocks A B C D).map f = from_blocks (A.map f) (B.map f) (C.map f) (D.map f) := @@ -456,6 +469,14 @@ end block_diag (block_diagonal M) = M := funext $ λ k, ext $ λ i j, block_diagonal_apply_eq M i j _ +lemma block_diagonal_injective [decidable_eq o] : + function.injective (block_diagonal : (o → matrix m n α) → matrix _ _ α) := +function.left_inverse.injective block_diag_block_diagonal + +@[simp] lemma block_diagonal_inj [decidable_eq o] {M N : o → matrix m n α} : + block_diagonal M = block_diagonal N ↔ M = N := +block_diagonal_injective.eq_iff + @[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 _ @@ -689,6 +710,14 @@ end block_diag' (block_diagonal' M) = M := funext $ λ k, ext $ λ i j, block_diagonal'_apply_eq M _ _ _ +lemma block_diagonal'_injective [decidable_eq o] : + function.injective (block_diagonal' : (Π i, matrix (m' i) (n' i) α) → matrix _ _ α) := +function.left_inverse.injective block_diag'_block_diagonal' + +@[simp] lemma block_diagonal'_inj [decidable_eq o] {M N : Π i, matrix (m' i) (n' i) α} : + block_diagonal' M = block_diagonal' N ↔ M = N := +block_diagonal'_injective.eq_iff + @[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 _ From fa2cb8a9e2b987db233e4e6eb47645feafba8861 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 26 Apr 2023 16:01:03 +0000 Subject: [PATCH 0879/1291] feat(data/set/pointwise/big_operators): image distributes across pointwise big operators (#18840) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The main results here are: * `set.image_finset_prod : f '' (∏ i in m, s i) = (∏ i in m, f '' s i)`, which says that the image under a monoid morphism commutes with the pointwise n-ary product of sets * `set.image_finset_prod_pi : (λ f : ι → α, ∏ i in l, f i) '' (l : set ι).pi S = (∏ i in l, S i)`, which says that turning a family of sets into a set of families and taking the product over each family is the same as taking the pointwise product. Both are n-ary versions of existing binary results. --- src/data/set/pointwise/big_operators.lean | 44 ++++++++++++++++++++--- 1 file changed, 40 insertions(+), 4 deletions(-) diff --git a/src/data/set/pointwise/big_operators.lean b/src/data/set/pointwise/big_operators.lean index 0345699d6e137..5460c2512846c 100644 --- a/src/data/set/pointwise/big_operators.lean +++ b/src/data/set/pointwise/big_operators.lean @@ -15,11 +15,35 @@ import data.set.pointwise.basic namespace set -section big_operators open_locale big_operators pointwise open function -variables {α : Type*} {ι : Type*} [comm_monoid α] +variables {ι α β F : Type*} + +section monoid +variables [monoid α] [monoid β] [monoid_hom_class F α β] + +@[to_additive] +lemma image_list_prod (f : F) : ∀ (l : list (set α)), + (f : α → β) '' l.prod = (l.map (λ s, f '' s)).prod +| [] := image_one.trans $ congr_arg singleton (map_one f) +| (a :: as) := by rw [list.map_cons, list.prod_cons, list.prod_cons, image_mul, image_list_prod] + +end monoid + +section comm_monoid +variables [comm_monoid α] [comm_monoid β] [monoid_hom_class F α β] + +@[to_additive] +lemma image_multiset_prod (f : F) : ∀ (m : multiset (set α)), + (f : α → β) '' m.prod = (m.map (λ s, f '' s)).prod := +quotient.ind $ by simpa only [multiset.quot_mk_to_coe, multiset.coe_prod, multiset.coe_map] + using image_list_prod f + +@[to_additive] +lemma image_finset_prod (f : F) (m : finset ι) (s : ι → set α) : + (f : α → β) '' (∏ i in m, s i) = (∏ i in m, f '' s i) := +(image_multiset_prod f _).trans $ congr_arg multiset.prod $ multiset.map_map _ _ _ /-- The n-ary version of `set.mem_mul`. -/ @[to_additive /-" The n-ary version of `set.mem_add`. "-/] @@ -127,8 +151,20 @@ lemma finset_prod_singleton {M ι : Type*} [comm_monoid M] (s : finset ι) (I : ∏ (i : ι) in s, ({I i} : set M) = {∏ (i : ι) in s, I i} := (map_prod (singleton_monoid_hom : M →* set M) _ _).symm -/-! TODO: define `decidable_mem_finset_prod` and `decidable_mem_finset_sum`. -/ +/-- The n-ary version of `set.image_mul_prod`. -/ +@[to_additive "The n-ary version of `set.add_image_prod`. "] +lemma image_finset_prod_pi (l : finset ι) (S : ι → set α) : + (λ f : ι → α, ∏ i in l, f i) '' (l : set ι).pi S = (∏ i in l, S i) := +by { ext, simp_rw [mem_finset_prod, mem_image, mem_pi, exists_prop, finset.mem_coe] } -end big_operators +/-- A special case of `set.image_finset_prod_pi` for `finset.univ`. -/ +@[to_additive "A special case of `set.image_finset_sum_pi` for `finset.univ`. "] +lemma image_fintype_prod_pi [fintype ι] (S : ι → set α) : + (λ f : ι → α, ∏ i, f i) '' univ.pi S = (∏ i, S i) := +by simpa only [finset.coe_univ] using image_finset_prod_pi finset.univ S + +end comm_monoid + +/-! TODO: define `decidable_mem_finset_prod` and `decidable_mem_finset_sum`. -/ end set From 178a32653e369dce2da68dc6b2694e385d484ef1 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Wed, 26 Apr 2023 16:01:04 +0000 Subject: [PATCH 0880/1291] chore(topology/category/Top/limits): split file (#18871) Per https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/!4.233487.20last.20minute.20split.3F This file is already being ported at https://github.com/leanprover-community/mathlib4/pull/3487, but: * it's not going so well right now * it is going well up to the point of the proposed new `limits/basic.lean` * that is sufficient to port the files needed for Copenhagen Co-authored-by: Scott Morrison --- .../calculating_colimits_in_Top.lean | 2 +- src/algebraic_geometry/open_immersion.lean | 1 + .../presheafed_space/has_colimits.lean | 2 +- .../fundamental_groupoid/product.lean | 2 +- src/algebraic_topology/simplicial_set.lean | 2 +- src/category_theory/cofiltered_system.lean | 2 +- src/category_theory/extensive.lean | 2 +- src/topology/category/CompHaus/basic.lean | 2 +- .../category/Profinite/cofiltered_limit.lean | 2 + src/topology/category/Top/limits.lean | 1049 ----------------- src/topology/category/Top/limits/basic.lean | 162 +++ .../category/Top/limits/cofiltered.lean | 121 ++ src/topology/category/Top/limits/konig.lean | 147 +++ .../category/Top/limits/products.lean | 311 +++++ .../category/Top/limits/pullbacks.lean | 400 +++++++ src/topology/gluing.lean | 2 +- 16 files changed, 1152 insertions(+), 1057 deletions(-) delete mode 100644 src/topology/category/Top/limits.lean create mode 100644 src/topology/category/Top/limits/basic.lean create mode 100644 src/topology/category/Top/limits/cofiltered.lean create mode 100644 src/topology/category/Top/limits/konig.lean create mode 100644 src/topology/category/Top/limits/products.lean create mode 100644 src/topology/category/Top/limits/pullbacks.lean diff --git a/docs/tutorial/category_theory/calculating_colimits_in_Top.lean b/docs/tutorial/category_theory/calculating_colimits_in_Top.lean index 9adee565f3fdd..5892030933eb5 100644 --- a/docs/tutorial/category_theory/calculating_colimits_in_Top.lean +++ b/docs/tutorial/category_theory/calculating_colimits_in_Top.lean @@ -1,4 +1,4 @@ -import topology.category.Top.limits +import topology.category.Top.limits.basic import topology.instances.real import topology.tactic diff --git a/src/algebraic_geometry/open_immersion.lean b/src/algebraic_geometry/open_immersion.lean index f14801af53854..66068eb2190b0 100644 --- a/src/algebraic_geometry/open_immersion.lean +++ b/src/algebraic_geometry/open_immersion.lean @@ -7,6 +7,7 @@ import algebraic_geometry.presheafed_space.has_colimits import category_theory.limits.shapes.binary_products import category_theory.limits.preserves.shapes.pullbacks import topology.sheaves.functors +import topology.category.Top.limits.pullbacks import algebraic_geometry.Scheme import category_theory.limits.shapes.strict_initial import category_theory.limits.shapes.comm_sq diff --git a/src/algebraic_geometry/presheafed_space/has_colimits.lean b/src/algebraic_geometry/presheafed_space/has_colimits.lean index eefd2a12e5c49..fe3ad8102f4d2 100644 --- a/src/algebraic_geometry/presheafed_space/has_colimits.lean +++ b/src/algebraic_geometry/presheafed_space/has_colimits.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ import algebraic_geometry.presheafed_space -import topology.category.Top.limits +import topology.category.Top.limits.basic import topology.sheaves.limits /-! diff --git a/src/algebraic_topology/fundamental_groupoid/product.lean b/src/algebraic_topology/fundamental_groupoid/product.lean index d3c3594941a83..8ecf0078ff45d 100644 --- a/src/algebraic_topology/fundamental_groupoid/product.lean +++ b/src/algebraic_topology/fundamental_groupoid/product.lean @@ -6,7 +6,7 @@ Authors: Praneeth Kolichala import category_theory.groupoid import algebraic_topology.fundamental_groupoid.basic -import topology.category.Top.limits +import topology.category.Top.limits.products import topology.homotopy.product /-! diff --git a/src/algebraic_topology/simplicial_set.lean b/src/algebraic_topology/simplicial_set.lean index 6727591420e4e..48a2fe3bac334 100644 --- a/src/algebraic_topology/simplicial_set.lean +++ b/src/algebraic_topology/simplicial_set.lean @@ -8,7 +8,7 @@ import algebraic_topology.topological_simplex import category_theory.limits.presheaf import category_theory.limits.types import category_theory.yoneda -import topology.category.Top.limits +import topology.category.Top.limits.basic /-! A simplicial set is just a simplicial object in `Type`, diff --git a/src/category_theory/cofiltered_system.lean b/src/category_theory/cofiltered_system.lean index 3ed2016a7f64b..73df182bb6ee8 100644 --- a/src/category_theory/cofiltered_system.lean +++ b/src/category_theory/cofiltered_system.lean @@ -5,7 +5,7 @@ Authors: Kyle Miller, Adam Topaz, Rémi Bottinelli, Junyan Xu -/ import category_theory.filtered import data.set.finite -import topology.category.Top.limits +import topology.category.Top.limits.konig /-! # Cofiltered systems diff --git a/src/category_theory/extensive.lean b/src/category_theory/extensive.lean index 1c4e5ba13f87c..6623631672aed 100644 --- a/src/category_theory/extensive.lean +++ b/src/category_theory/extensive.lean @@ -6,7 +6,7 @@ Authors: Andrew Yang import category_theory.limits.shapes.comm_sq import category_theory.limits.shapes.strict_initial import category_theory.limits.shapes.types -import topology.category.Top.limits +import topology.category.Top.limits.pullbacks import category_theory.limits.functor_category /-! diff --git a/src/topology/category/CompHaus/basic.lean b/src/topology/category/CompHaus/basic.lean index 2c7fa0f85ad20..287dc5ac0ea7b 100644 --- a/src/topology/category/CompHaus/basic.lean +++ b/src/topology/category/CompHaus/basic.lean @@ -8,7 +8,7 @@ import category_theory.adjunction.reflective import topology.stone_cech import category_theory.monad.limits import topology.urysohns_lemma -import topology.category.Top.limits +import topology.category.Top.limits.basic /-! # The category of Compact Hausdorff Spaces diff --git a/src/topology/category/Profinite/cofiltered_limit.lean b/src/topology/category/Profinite/cofiltered_limit.lean index 001599ef16aa5..1fb982d1ad7d3 100644 --- a/src/topology/category/Profinite/cofiltered_limit.lean +++ b/src/topology/category/Profinite/cofiltered_limit.lean @@ -6,6 +6,8 @@ Authors: Adam Topaz import topology.category.Profinite.basic import topology.locally_constant.basic import topology.discrete_quotient +import topology.category.Top.limits.cofiltered +import topology.category.Top.limits.konig /-! # Cofiltered limits of profinite sets. diff --git a/src/topology/category/Top/limits.lean b/src/topology/category/Top/limits.lean deleted file mode 100644 index 4069e38aae9db..0000000000000 --- a/src/topology/category/Top/limits.lean +++ /dev/null @@ -1,1049 +0,0 @@ -/- -Copyright (c) 2017 Scott Morrison. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Patrick Massot, Scott Morrison, Mario Carneiro, Andrew Yang --/ -import topology.category.Top.epi_mono -import category_theory.category.ulift -import category_theory.limits.concrete_category -import category_theory.concrete_category.elementwise - -/-! -# The category of topological spaces has all limits and colimits - -Further, these limits and colimits are preserved by the forgetful functor --- that is, the -underlying types are just the limits in the category of types. --/ - -open topological_space -open category_theory -open category_theory.limits -open opposite - -universes u v w - -noncomputable theory - -namespace Top - -variables {J : Type v} [small_category J] - -local notation `forget` := forget Top - -/-- -A choice of limit cone for a functor `F : J ⥤ Top`. -Generally you should just use `limit.cone F`, unless you need the actual definition -(which is in terms of `types.limit_cone`). --/ -def limit_cone (F : J ⥤ Top.{max v u}) : cone F := -{ X := Top.of {u : Π j : J, F.obj j | ∀ {i j : J} (f : i ⟶ j), F.map f (u i) = u j}, - π := - { app := λ j, - { to_fun := λ u, u.val j, - continuous_to_fun := show continuous ((λ u : Π j : J, F.obj j, u j) ∘ subtype.val), - by continuity } } } - -/-- -A choice of limit cone for a functor `F : J ⥤ Top` whose topology is defined as an -infimum of topologies infimum. -Generally you should just use `limit.cone F`, unless you need the actual definition -(which is in terms of `types.limit_cone`). --/ -def limit_cone_infi (F : J ⥤ Top.{max v u}) : cone F := -{ X := ⟨(types.limit_cone (F ⋙ forget)).X, ⨅j, - (F.obj j).str.induced ((types.limit_cone (F ⋙ forget)).π.app j)⟩, - π := - { app := λ j, ⟨(types.limit_cone (F ⋙ forget)).π.app j, - continuous_iff_le_induced.mpr (infi_le _ _)⟩, - naturality' := λ j j' f, continuous_map.coe_injective - ((types.limit_cone (F ⋙ forget)).π.naturality f) } } - -/-- -The chosen cone `Top.limit_cone F` for a functor `F : J ⥤ Top` is a limit cone. -Generally you should just use `limit.is_limit F`, unless you need the actual definition -(which is in terms of `types.limit_cone_is_limit`). --/ -def limit_cone_is_limit (F : J ⥤ Top.{max v u}) : is_limit (limit_cone F) := -{ lift := λ S, { to_fun := λ x, ⟨λ j, S.π.app _ x, λ i j f, by { dsimp, erw ← S.w f, refl }⟩ }, - uniq' := λ S m h, by { ext : 3, simpa [← h] } } - -/-- -The chosen cone `Top.limit_cone_infi F` for a functor `F : J ⥤ Top` is a limit cone. -Generally you should just use `limit.is_limit F`, unless you need the actual definition -(which is in terms of `types.limit_cone_is_limit`). --/ -def limit_cone_infi_is_limit (F : J ⥤ Top.{max v u}) : is_limit (limit_cone_infi F) := -by { refine is_limit.of_faithful forget (types.limit_cone_is_limit _) (λ s, ⟨_, _⟩) (λ s, rfl), - exact continuous_iff_coinduced_le.mpr (le_infi $ λ j, - coinduced_le_iff_le_induced.mp $ (continuous_iff_coinduced_le.mp (s.π.app j).continuous : - _) ) } - -instance Top_has_limits_of_size : has_limits_of_size.{v} Top.{max v u} := -{ has_limits_of_shape := λ J 𝒥, by exactI - { has_limit := λ F, has_limit.mk { cone := limit_cone F, is_limit := limit_cone_is_limit F } } } - -instance Top_has_limits : has_limits Top.{u} := Top.Top_has_limits_of_size.{u u} - -instance forget_preserves_limits_of_size : - preserves_limits_of_size.{v v} (forget : Top.{max v u} ⥤ Type (max v u)) := -{ preserves_limits_of_shape := λ J 𝒥, - { preserves_limit := λ F, - by exactI preserves_limit_of_preserves_limit_cone - (limit_cone_is_limit F) (types.limit_cone_is_limit (F ⋙ forget)) } } - -instance forget_preserves_limits : preserves_limits (forget : Top.{u} ⥤ Type u) := -Top.forget_preserves_limits_of_size.{u u} - -/-- -A choice of colimit cocone for a functor `F : J ⥤ Top`. -Generally you should just use `colimit.coone F`, unless you need the actual definition -(which is in terms of `types.colimit_cocone`). --/ -def colimit_cocone (F : J ⥤ Top.{max v u}) : cocone F := -{ X := ⟨(types.colimit_cocone (F ⋙ forget)).X, ⨆ j, - (F.obj j).str.coinduced ((types.colimit_cocone (F ⋙ forget)).ι.app j)⟩, - ι := - { app := λ j, ⟨(types.colimit_cocone (F ⋙ forget)).ι.app j, - continuous_iff_coinduced_le.mpr (le_supr _ j)⟩, - naturality' := λ j j' f, continuous_map.coe_injective - ((types.colimit_cocone (F ⋙ forget)).ι.naturality f) } } - -/-- -The chosen cocone `Top.colimit_cocone F` for a functor `F : J ⥤ Top` is a colimit cocone. -Generally you should just use `colimit.is_colimit F`, unless you need the actual definition -(which is in terms of `types.colimit_cocone_is_colimit`). --/ -def colimit_cocone_is_colimit (F : J ⥤ Top.{max v u}) : is_colimit (colimit_cocone F) := -by { refine is_colimit.of_faithful forget (types.colimit_cocone_is_colimit _) (λ s, ⟨_, _⟩) - (λ s, rfl), - exact continuous_iff_le_induced.mpr (supr_le $ λ j, - coinduced_le_iff_le_induced.mp $ (continuous_iff_coinduced_le.mp (s.ι.app j).continuous : - _) ) } - -instance Top_has_colimits_of_size : has_colimits_of_size.{v} Top.{max v u} := -{ has_colimits_of_shape := λ J 𝒥, by exactI - { has_colimit := λ F, has_colimit.mk { cocone := colimit_cocone F, is_colimit := - colimit_cocone_is_colimit F } } } - -instance Top_has_colimits : has_colimits Top.{u} := Top.Top_has_colimits_of_size.{u u} - -instance forget_preserves_colimits_of_size : - preserves_colimits_of_size.{v v} (forget : Top.{max v u} ⥤ Type (max v u)) := -{ preserves_colimits_of_shape := λ J 𝒥, - { preserves_colimit := λ F, - by exactI preserves_colimit_of_preserves_colimit_cocone - (colimit_cocone_is_colimit F) (types.colimit_cocone_is_colimit (F ⋙ forget)) } } - -instance forget_preserves_colimits : preserves_colimits (forget : Top.{u} ⥤ Type u) := -Top.forget_preserves_colimits_of_size.{u u} - -/-- The projection from the product as a bundled continous map. -/ -abbreviation pi_π {ι : Type v} (α : ι → Top.{max v u}) (i : ι) : Top.of (Π i, α i) ⟶ α i := -⟨λ f, f i, continuous_apply i⟩ - -/-- The explicit fan of a family of topological spaces given by the pi type. -/ -@[simps X π_app] -def pi_fan {ι : Type v} (α : ι → Top.{max v u}) : fan α := -fan.mk (Top.of (Π i, α i)) (pi_π α) - -/-- The constructed fan is indeed a limit -/ -def pi_fan_is_limit {ι : Type v} (α : ι → Top.{max v u}) : is_limit (pi_fan α) := -{ lift := λ S, { to_fun := λ s i, S.π.app ⟨i⟩ s }, - uniq' := by { intros S m h, ext x i, simp [← h ⟨i⟩] }, - fac' := λ s j, by { cases j, tidy, }, } - -/-- -The product is homeomorphic to the product of the underlying spaces, -equipped with the product topology. --/ -def pi_iso_pi {ι : Type v} (α : ι → Top.{max v u}) : ∏ α ≅ Top.of (Π i, α i) := -(limit.is_limit _).cone_point_unique_up_to_iso (pi_fan_is_limit α) - -@[simp, reassoc] -lemma pi_iso_pi_inv_π {ι : Type v} (α : ι → Top.{max v u}) (i : ι) : - (pi_iso_pi α).inv ≫ pi.π α i = pi_π α i := -by simp [pi_iso_pi] - -@[simp] -lemma pi_iso_pi_inv_π_apply {ι : Type v} (α : ι → Top.{max v u}) (i : ι) (x : Π i, α i) : - (pi.π α i : _) ((pi_iso_pi α).inv x) = x i := -concrete_category.congr_hom (pi_iso_pi_inv_π α i) x - -@[simp] -lemma pi_iso_pi_hom_apply {ι : Type v} (α : ι → Top.{max v u}) (i : ι) (x : ∏ α) : - (pi_iso_pi α).hom x i = (pi.π α i : _) x := -begin - have := pi_iso_pi_inv_π α i, - rw iso.inv_comp_eq at this, - exact concrete_category.congr_hom this x -end - -/-- The inclusion to the coproduct as a bundled continous map. -/ -abbreviation sigma_ι {ι : Type v} (α : ι → Top.{max v u}) (i : ι) : α i ⟶ Top.of (Σ i, α i) := -⟨sigma.mk i⟩ - -/-- The explicit cofan of a family of topological spaces given by the sigma type. -/ -@[simps X ι_app] -def sigma_cofan {ι : Type v} (α : ι → Top.{max v u}) : cofan α := -cofan.mk (Top.of (Σ i, α i)) (sigma_ι α) - -/-- The constructed cofan is indeed a colimit -/ -def sigma_cofan_is_colimit {ι : Type v} (α : ι → Top.{max v u}) : is_colimit (sigma_cofan α) := -{ desc := λ S, { to_fun := λ s, S.ι.app ⟨s.1⟩ s.2, - continuous_to_fun := continuous_sigma $ λ i, map_continuous (S.ι.app ⟨i⟩) }, - uniq' := by { intros S m h, ext ⟨i, x⟩, simp [← h ⟨i⟩] }, - fac' := λ s j, by { cases j, tidy, }, } - -/-- -The coproduct is homeomorphic to the disjoint union of the topological spaces. --/ -def sigma_iso_sigma {ι : Type v} (α : ι → Top.{max v u}) : ∐ α ≅ Top.of (Σ i, α i) := -(colimit.is_colimit _).cocone_point_unique_up_to_iso (sigma_cofan_is_colimit α) - -@[simp, reassoc] -lemma sigma_iso_sigma_hom_ι {ι : Type v} (α : ι → Top.{max v u}) (i : ι) : - sigma.ι α i ≫ (sigma_iso_sigma α).hom = sigma_ι α i := -by simp [sigma_iso_sigma] - -@[simp] -lemma sigma_iso_sigma_hom_ι_apply {ι : Type v} (α : ι → Top.{max v u}) (i : ι) (x : α i) : - (sigma_iso_sigma α).hom ((sigma.ι α i : _) x) = sigma.mk i x := -concrete_category.congr_hom (sigma_iso_sigma_hom_ι α i) x - -@[simp] -lemma sigma_iso_sigma_inv_apply {ι : Type v} (α : ι → Top.{max v u}) (i : ι) (x : α i) : - (sigma_iso_sigma α).inv ⟨i, x⟩ = (sigma.ι α i : _) x := -by { rw [← sigma_iso_sigma_hom_ι_apply, ← comp_app], simp, } - -lemma induced_of_is_limit {F : J ⥤ Top.{max v u}} (C : cone F) (hC : is_limit C) : - C.X.topological_space = ⨅ j, (F.obj j).topological_space.induced (C.π.app j) := -begin - let homeo := homeo_of_iso (hC.cone_point_unique_up_to_iso (limit_cone_infi_is_limit F)), - refine homeo.inducing.induced.trans _, - change induced homeo (⨅ (j : J), _) = _, - simpa [induced_infi, induced_compose], -end - -lemma limit_topology (F : J ⥤ Top.{max v u}) : - (limit F).topological_space = ⨅ j, (F.obj j).topological_space.induced (limit.π F j) := -induced_of_is_limit _ (limit.is_limit F) - -section prod - -/-- The first projection from the product. -/ -abbreviation prod_fst {X Y : Top.{u}} : Top.of (X × Y) ⟶ X := ⟨prod.fst⟩ - -/-- The second projection from the product. -/ -abbreviation prod_snd {X Y : Top.{u}} : Top.of (X × Y) ⟶ Y := ⟨prod.snd⟩ - -/-- The explicit binary cofan of `X, Y` given by `X × Y`. -/ -def prod_binary_fan (X Y : Top.{u}) : binary_fan X Y := -binary_fan.mk prod_fst prod_snd - -/-- The constructed binary fan is indeed a limit -/ -def prod_binary_fan_is_limit (X Y : Top.{u}) : is_limit (prod_binary_fan X Y) := -{ lift := λ (S : binary_fan X Y), { to_fun := λ s, (S.fst s, S.snd s) }, - fac' := begin - rintros S (_|_), - tidy - end, - uniq' := begin - intros S m h, - ext x, - { specialize h ⟨walking_pair.left⟩, - apply_fun (λ e, (e x)) at h, - exact h }, - { specialize h ⟨walking_pair.right⟩, - apply_fun (λ e, (e x)) at h, - exact h }, - end } - -/-- -The homeomorphism between `X ⨯ Y` and the set-theoretic product of `X` and `Y`, -equipped with the product topology. --/ -def prod_iso_prod (X Y : Top.{u}) : X ⨯ Y ≅ Top.of (X × Y) := -(limit.is_limit _).cone_point_unique_up_to_iso (prod_binary_fan_is_limit X Y) - -@[simp, reassoc] lemma prod_iso_prod_hom_fst (X Y : Top.{u}) : - (prod_iso_prod X Y).hom ≫ prod_fst = limits.prod.fst := -by simpa [← iso.eq_inv_comp, prod_iso_prod] - -@[simp, reassoc] lemma prod_iso_prod_hom_snd (X Y : Top.{u}) : - (prod_iso_prod X Y).hom ≫ prod_snd = limits.prod.snd := -by simpa [← iso.eq_inv_comp, prod_iso_prod] - -@[simp] lemma prod_iso_prod_hom_apply {X Y : Top.{u}} (x : X ⨯ Y) : - (prod_iso_prod X Y).hom x = - ((limits.prod.fst : X ⨯ Y ⟶ _) x, (limits.prod.snd : X ⨯ Y ⟶ _) x) := -begin - ext, - { exact concrete_category.congr_hom (prod_iso_prod_hom_fst X Y) x }, - { exact concrete_category.congr_hom (prod_iso_prod_hom_snd X Y) x } -end - -@[simp, reassoc, elementwise] lemma prod_iso_prod_inv_fst (X Y : Top.{u}) : - (prod_iso_prod X Y).inv ≫ limits.prod.fst = prod_fst := -by simp [iso.inv_comp_eq] - -@[simp, reassoc, elementwise] lemma prod_iso_prod_inv_snd (X Y : Top.{u}) : - (prod_iso_prod X Y).inv ≫ limits.prod.snd = prod_snd := -by simp [iso.inv_comp_eq] - -lemma prod_topology {X Y : Top} : - (X ⨯ Y).topological_space = - induced (limits.prod.fst : X ⨯ Y ⟶ _) X.topological_space ⊓ - induced (limits.prod.snd : X ⨯ Y ⟶ _) Y.topological_space := -begin - let homeo := homeo_of_iso (prod_iso_prod X Y), - refine homeo.inducing.induced.trans _, - change induced homeo (_ ⊓ _) = _, - simpa [induced_compose] -end - -lemma range_prod_map {W X Y Z : Top.{u}} (f : W ⟶ Y) (g : X ⟶ Z) : - set.range (limits.prod.map f g) = - (limits.prod.fst : Y ⨯ Z ⟶ _) ⁻¹' (set.range f) ∩ - (limits.prod.snd : Y ⨯ Z ⟶ _) ⁻¹' (set.range g) := -begin - ext, - split, - { rintros ⟨y, rfl⟩, - simp only [set.mem_preimage, set.mem_range, set.mem_inter_iff, ←comp_apply], - simp only [limits.prod.map_fst, limits.prod.map_snd, - exists_apply_eq_apply, comp_apply, and_self] }, - { rintros ⟨⟨x₁, hx₁⟩, ⟨x₂, hx₂⟩⟩, - use (prod_iso_prod W X).inv (x₁, x₂), - apply concrete.limit_ext, - rintro ⟨⟨⟩⟩, - { simp only [← comp_apply, category.assoc], erw limits.prod.map_fst, simp [hx₁] }, - { simp only [← comp_apply, category.assoc], erw limits.prod.map_snd, simp [hx₂] } } -end - -lemma inducing_prod_map {W X Y Z : Top} {f : W ⟶ X} {g : Y ⟶ Z} - (hf : inducing f) (hg : inducing g) : inducing (limits.prod.map f g) := -begin - constructor, - simp only [prod_topology, induced_compose, ←coe_comp, limits.prod.map_fst, limits.prod.map_snd, - induced_inf], - simp only [coe_comp], - rw [← @induced_compose _ _ _ _ _ f, ← @induced_compose _ _ _ _ _ g, ← hf.induced, ← hg.induced] -end - -lemma embedding_prod_map {W X Y Z : Top} {f : W ⟶ X} {g : Y ⟶ Z} - (hf : embedding f) (hg : embedding g) : embedding (limits.prod.map f g) := -⟨inducing_prod_map hf.to_inducing hg.to_inducing, -begin - haveI := (Top.mono_iff_injective _).mpr hf.inj, - haveI := (Top.mono_iff_injective _).mpr hg.inj, - exact (Top.mono_iff_injective _).mp infer_instance -end⟩ - -end prod - -section pullback - -variables {X Y Z : Top.{u}} - -/-- The first projection from the pullback. -/ -abbreviation pullback_fst (f : X ⟶ Z) (g : Y ⟶ Z) : Top.of { p : X × Y // f p.1 = g p.2 } ⟶ X := -⟨prod.fst ∘ subtype.val⟩ - -/-- The second projection from the pullback. -/ -abbreviation pullback_snd (f : X ⟶ Z) (g : Y ⟶ Z) : Top.of { p : X × Y // f p.1 = g p.2 } ⟶ Y := -⟨prod.snd ∘ subtype.val⟩ - -/-- The explicit pullback cone of `X, Y` given by `{ p : X × Y // f p.1 = g p.2 }`. -/ -def pullback_cone (f : X ⟶ Z) (g : Y ⟶ Z) : pullback_cone f g := -pullback_cone.mk (pullback_fst f g) (pullback_snd f g) (by { ext ⟨x, h⟩, simp [h] }) - -/-- The constructed cone is a limit. -/ -def pullback_cone_is_limit (f : X ⟶ Z) (g : Y ⟶ Z) : - is_limit (pullback_cone f g) := pullback_cone.is_limit_aux' _ -begin - intro s, - split, swap, - exact { to_fun := λ x, ⟨⟨s.fst x, s.snd x⟩, - by simpa using concrete_category.congr_hom s.condition x⟩ }, - refine ⟨_,_,_⟩, - { ext, delta pullback_cone, simp }, - { ext, delta pullback_cone, simp }, - { intros m h₁ h₂, - ext x, - { simpa using concrete_category.congr_hom h₁ x }, - { simpa using concrete_category.congr_hom h₂ x } } -end - -/-- The pullback of two maps can be identified as a subspace of `X × Y`. -/ -def pullback_iso_prod_subtype (f : X ⟶ Z) (g : Y ⟶ Z) : - pullback f g ≅ Top.of { p : X × Y // f p.1 = g p.2 } := -(limit.is_limit _).cone_point_unique_up_to_iso (pullback_cone_is_limit f g) - -@[simp, reassoc] lemma pullback_iso_prod_subtype_inv_fst (f : X ⟶ Z) (g : Y ⟶ Z) : - (pullback_iso_prod_subtype f g).inv ≫ pullback.fst = pullback_fst f g := -by simpa [pullback_iso_prod_subtype] - -@[simp] lemma pullback_iso_prod_subtype_inv_fst_apply (f : X ⟶ Z) (g : Y ⟶ Z) - (x : { p : X × Y // f p.1 = g p.2 }) : - (pullback.fst : pullback f g ⟶ _) ((pullback_iso_prod_subtype f g).inv x) = (x : X × Y).fst := -concrete_category.congr_hom (pullback_iso_prod_subtype_inv_fst f g) x - -@[simp, reassoc] lemma pullback_iso_prod_subtype_inv_snd (f : X ⟶ Z) (g : Y ⟶ Z) : - (pullback_iso_prod_subtype f g).inv ≫ pullback.snd = pullback_snd f g := -by simpa [pullback_iso_prod_subtype] - -@[simp] lemma pullback_iso_prod_subtype_inv_snd_apply (f : X ⟶ Z) (g : Y ⟶ Z) - (x : { p : X × Y // f p.1 = g p.2 }) : - (pullback.snd : pullback f g ⟶ _) ((pullback_iso_prod_subtype f g).inv x) = (x : X × Y).snd := -concrete_category.congr_hom (pullback_iso_prod_subtype_inv_snd f g) x - -lemma pullback_iso_prod_subtype_hom_fst (f : X ⟶ Z) (g : Y ⟶ Z) : - (pullback_iso_prod_subtype f g).hom ≫ pullback_fst f g = pullback.fst := -by rw [←iso.eq_inv_comp, pullback_iso_prod_subtype_inv_fst] - -lemma pullback_iso_prod_subtype_hom_snd (f : X ⟶ Z) (g : Y ⟶ Z) : - (pullback_iso_prod_subtype f g).hom ≫ pullback_snd f g = pullback.snd := -by rw [←iso.eq_inv_comp, pullback_iso_prod_subtype_inv_snd] - -@[simp] lemma pullback_iso_prod_subtype_hom_apply {f : X ⟶ Z} {g : Y ⟶ Z} - (x : pullback f g) : (pullback_iso_prod_subtype f g).hom x = - ⟨⟨(pullback.fst : pullback f g ⟶ _) x, (pullback.snd : pullback f g ⟶ _) x⟩, - by simpa using concrete_category.congr_hom pullback.condition x⟩ := -begin - ext, - exacts [concrete_category.congr_hom (pullback_iso_prod_subtype_hom_fst f g) x, - concrete_category.congr_hom (pullback_iso_prod_subtype_hom_snd f g) x] -end - -lemma pullback_topology {X Y Z : Top.{u}} (f : X ⟶ Z) (g : Y ⟶ Z) : - (pullback f g).topological_space = - induced (pullback.fst : pullback f g ⟶ _) X.topological_space ⊓ - induced (pullback.snd : pullback f g ⟶ _) Y.topological_space := -begin - let homeo := homeo_of_iso (pullback_iso_prod_subtype f g), - refine homeo.inducing.induced.trans _, - change induced homeo (induced _ (_ ⊓ _)) = _, - simpa [induced_compose] -end - -lemma range_pullback_to_prod {X Y Z : Top} (f : X ⟶ Z) (g : Y ⟶ Z) : - set.range (prod.lift pullback.fst pullback.snd : pullback f g ⟶ X ⨯ Y) = - { x | (limits.prod.fst ≫ f) x = (limits.prod.snd ≫ g) x } := -begin - ext x, - split, - { rintros ⟨y, rfl⟩, - simp only [←comp_apply, set.mem_set_of_eq], - congr' 1, - simp [pullback.condition] }, - { intro h, - use (pullback_iso_prod_subtype f g).inv ⟨⟨_, _⟩, h⟩, - apply concrete.limit_ext, - rintro ⟨⟨⟩⟩; simp, } -end - -lemma inducing_pullback_to_prod {X Y Z : Top} (f : X ⟶ Z) (g : Y ⟶ Z) : - inducing ⇑(prod.lift pullback.fst pullback.snd : pullback f g ⟶ X ⨯ Y) := -⟨by simp [prod_topology, pullback_topology, induced_compose, ←coe_comp]⟩ - -lemma embedding_pullback_to_prod {X Y Z : Top} (f : X ⟶ Z) (g : Y ⟶ Z) : - embedding ⇑(prod.lift pullback.fst pullback.snd : pullback f g ⟶ X ⨯ Y) := -⟨inducing_pullback_to_prod f g, (Top.mono_iff_injective _).mp infer_instance⟩ - -/-- If the map `S ⟶ T` is mono, then there is a description of the image of `W ×ₛ X ⟶ Y ×ₜ Z`. -/ -lemma range_pullback_map {W X Y Z S T : Top} (f₁ : W ⟶ S) (f₂ : X ⟶ S) - (g₁ : Y ⟶ T) (g₂ : Z ⟶ T) (i₁ : W ⟶ Y) (i₂ : X ⟶ Z) (i₃ : S ⟶ T) [H₃ : mono i₃] - (eq₁ : f₁ ≫ i₃ = i₁ ≫ g₁) (eq₂ : f₂ ≫ i₃ = i₂ ≫ g₂) : - set.range (pullback.map f₁ f₂ g₁ g₂ i₁ i₂ i₃ eq₁ eq₂) = - (pullback.fst : pullback g₁ g₂ ⟶ _) ⁻¹' (set.range i₁) ∩ - (pullback.snd : pullback g₁ g₂ ⟶ _) ⁻¹' (set.range i₂) := -begin - ext, - split, - { rintro ⟨y, rfl⟩, simp, }, - rintros ⟨⟨x₁, hx₁⟩, ⟨x₂, hx₂⟩⟩, - have : f₁ x₁ = f₂ x₂, - { apply (Top.mono_iff_injective _).mp H₃, - simp only [←comp_apply, eq₁, eq₂], - simp only [comp_apply, hx₁, hx₂], - simp only [←comp_apply, pullback.condition] }, - use (pullback_iso_prod_subtype f₁ f₂).inv ⟨⟨x₁, x₂⟩, this⟩, - apply concrete.limit_ext, - rintros (_|_|_), - { simp only [Top.comp_app, limit.lift_π_apply, category.assoc, pullback_cone.mk_π_app_one, - hx₁, pullback_iso_prod_subtype_inv_fst_apply, subtype.coe_mk], - simp only [← comp_apply], - congr, - apply limit.w _ walking_cospan.hom.inl }, - { simp [hx₁] }, - { simp [hx₂] }, -end - -lemma pullback_fst_range {X Y S : Top} (f : X ⟶ S) (g : Y ⟶ S) : - set.range (pullback.fst : pullback f g ⟶ _) = { x : X | ∃ y : Y, f x = g y} := -begin - ext x, - split, - { rintro ⟨y, rfl⟩, - use (pullback.snd : pullback f g ⟶ _) y, - exact concrete_category.congr_hom pullback.condition y }, - { rintro ⟨y, eq⟩, - use (Top.pullback_iso_prod_subtype f g).inv ⟨⟨x, y⟩, eq⟩, - simp }, -end - -lemma pullback_snd_range {X Y S : Top} (f : X ⟶ S) (g : Y ⟶ S) : - set.range (pullback.snd : pullback f g ⟶ _) = { y : Y | ∃ x : X, f x = g y} := -begin - ext y, - split, - { rintro ⟨x, rfl⟩, - use (pullback.fst : pullback f g ⟶ _) x, - exact concrete_category.congr_hom pullback.condition x }, - { rintro ⟨x, eq⟩, - use (Top.pullback_iso_prod_subtype f g).inv ⟨⟨x, y⟩, eq⟩, - simp }, -end - -/-- -If there is a diagram where the morphisms `W ⟶ Y` and `X ⟶ Z` are embeddings, -then the induced morphism `W ×ₛ X ⟶ Y ×ₜ Z` is also an embedding. - - W ⟶ Y - ↘ ↘ - S ⟶ T - ↗ ↗ - X ⟶ Z --/ -lemma pullback_map_embedding_of_embeddings {W X Y Z S T : Top} - (f₁ : W ⟶ S) (f₂ : X ⟶ S) (g₁ : Y ⟶ T) (g₂ : Z ⟶ T) {i₁ : W ⟶ Y} {i₂ : X ⟶ Z} - (H₁ : embedding i₁) (H₂ : embedding i₂) (i₃ : S ⟶ T) - (eq₁ : f₁ ≫ i₃ = i₁ ≫ g₁) (eq₂ : f₂ ≫ i₃ = i₂ ≫ g₂) : - embedding (pullback.map f₁ f₂ g₁ g₂ i₁ i₂ i₃ eq₁ eq₂) := -begin - refine embedding_of_embedding_compose (continuous_map.continuous_to_fun _) - (show continuous (prod.lift pullback.fst pullback.snd : pullback g₁ g₂ ⟶ Y ⨯ Z), from - continuous_map.continuous_to_fun _) _, - suffices : embedding - (prod.lift pullback.fst pullback.snd ≫ limits.prod.map i₁ i₂ : pullback f₁ f₂ ⟶ _), - { simpa [←coe_comp] using this }, - rw coe_comp, - refine embedding.comp (embedding_prod_map H₁ H₂) - (embedding_pullback_to_prod _ _) -end - -/-- -If there is a diagram where the morphisms `W ⟶ Y` and `X ⟶ Z` are open embeddings, and `S ⟶ T` -is mono, then the induced morphism `W ×ₛ X ⟶ Y ×ₜ Z` is also an open embedding. - W ⟶ Y - ↘ ↘ - S ⟶ T - ↗ ↗ - X ⟶ Z --/ -lemma pullback_map_open_embedding_of_open_embeddings {W X Y Z S T : Top} - (f₁ : W ⟶ S) (f₂ : X ⟶ S) (g₁ : Y ⟶ T) (g₂ : Z ⟶ T) {i₁ : W ⟶ Y} {i₂ : X ⟶ Z} - (H₁ : open_embedding i₁) (H₂ : open_embedding i₂) (i₃ : S ⟶ T) [H₃ : mono i₃] - (eq₁ : f₁ ≫ i₃ = i₁ ≫ g₁) (eq₂ : f₂ ≫ i₃ = i₂ ≫ g₂) : - open_embedding (pullback.map f₁ f₂ g₁ g₂ i₁ i₂ i₃ eq₁ eq₂) := -begin - split, - { apply pullback_map_embedding_of_embeddings - f₁ f₂ g₁ g₂ H₁.to_embedding H₂.to_embedding i₃ eq₁ eq₂ }, - { rw range_pullback_map, - apply is_open.inter; apply continuous.is_open_preimage, - continuity, - exacts [H₁.open_range, H₂.open_range] } -end - -lemma snd_embedding_of_left_embedding {X Y S : Top} - {f : X ⟶ S} (H : embedding f) (g : Y ⟶ S) : - embedding ⇑(pullback.snd : pullback f g ⟶ Y) := -begin - convert (homeo_of_iso (as_iso (pullback.snd : pullback (𝟙 S) g ⟶ _))).embedding.comp - (pullback_map_embedding_of_embeddings f g (𝟙 _) g H - (homeo_of_iso (iso.refl _)).embedding (𝟙 _) rfl (by simp)), - erw ←coe_comp, - simp -end - -lemma fst_embedding_of_right_embedding {X Y S : Top} - (f : X ⟶ S) {g : Y ⟶ S} (H : embedding g) : - embedding ⇑(pullback.fst : pullback f g ⟶ X) := -begin - convert (homeo_of_iso (as_iso (pullback.fst : pullback f (𝟙 S) ⟶ _))).embedding.comp - (pullback_map_embedding_of_embeddings f g f (𝟙 _) - (homeo_of_iso (iso.refl _)).embedding H (𝟙 _) rfl (by simp)), - erw ←coe_comp, - simp -end - -lemma embedding_of_pullback_embeddings {X Y S : Top} - {f : X ⟶ S} {g : Y ⟶ S} (H₁ : embedding f) (H₂ : embedding g) : - embedding (limit.π (cospan f g) walking_cospan.one) := -begin - convert H₂.comp (snd_embedding_of_left_embedding H₁ g), - erw ←coe_comp, - congr, - exact (limit.w _ walking_cospan.hom.inr).symm -end - -lemma snd_open_embedding_of_left_open_embedding {X Y S : Top} - {f : X ⟶ S} (H : open_embedding f) (g : Y ⟶ S) : - open_embedding ⇑(pullback.snd : pullback f g ⟶ Y) := -begin - convert (homeo_of_iso (as_iso (pullback.snd : pullback (𝟙 S) g ⟶ _))).open_embedding.comp - (pullback_map_open_embedding_of_open_embeddings f g (𝟙 _) g H - (homeo_of_iso (iso.refl _)).open_embedding (𝟙 _) rfl (by simp)), - erw ←coe_comp, - simp -end - -lemma fst_open_embedding_of_right_open_embedding {X Y S : Top} - (f : X ⟶ S) {g : Y ⟶ S} (H : open_embedding g) : - open_embedding ⇑(pullback.fst : pullback f g ⟶ X) := -begin - convert (homeo_of_iso (as_iso (pullback.fst : pullback f (𝟙 S) ⟶ _))).open_embedding.comp - (pullback_map_open_embedding_of_open_embeddings f g f (𝟙 _) - (homeo_of_iso (iso.refl _)).open_embedding H (𝟙 _) rfl (by simp)), - erw ←coe_comp, - simp -end - -/-- If `X ⟶ S`, `Y ⟶ S` are open embeddings, then so is `X ×ₛ Y ⟶ S`. -/ -lemma open_embedding_of_pullback_open_embeddings {X Y S : Top} - {f : X ⟶ S} {g : Y ⟶ S} (H₁ : open_embedding f) (H₂ : open_embedding g) : - open_embedding (limit.π (cospan f g) walking_cospan.one) := -begin - convert H₂.comp (snd_open_embedding_of_left_open_embedding H₁ g), - erw ←coe_comp, - congr, - exact (limit.w _ walking_cospan.hom.inr).symm -end - -lemma fst_iso_of_right_embedding_range_subset {X Y S : Top} (f : X ⟶ S) {g : Y ⟶ S} - (hg : embedding g) (H : set.range f ⊆ set.range g) : is_iso (pullback.fst : pullback f g ⟶ X) := -begin - let : (pullback f g : Top) ≃ₜ X := - (homeomorph.of_embedding _ (fst_embedding_of_right_embedding f hg)).trans - { to_fun := coe, - inv_fun := (λ x, ⟨x, - by { rw pullback_fst_range, exact ⟨_, (H (set.mem_range_self x)).some_spec.symm⟩ }⟩), - left_inv := λ ⟨_,_⟩, rfl, - right_inv := λ x, rfl }, - convert is_iso.of_iso (iso_of_homeo this), - ext, - refl -end - -lemma snd_iso_of_left_embedding_range_subset {X Y S : Top} {f : X ⟶ S} (hf : embedding f) - (g : Y ⟶ S) (H : set.range g ⊆ set.range f) : is_iso (pullback.snd : pullback f g ⟶ Y) := -begin - let : (pullback f g : Top) ≃ₜ Y := - (homeomorph.of_embedding _ (snd_embedding_of_left_embedding hf g)).trans - { to_fun := coe, - inv_fun := (λ x, ⟨x, - by { rw pullback_snd_range, exact ⟨_, (H (set.mem_range_self x)).some_spec⟩ }⟩), - left_inv := λ ⟨_,_⟩, rfl, - right_inv := λ x, rfl }, - convert is_iso.of_iso (iso_of_homeo this), - ext, - refl -end - -lemma pullback_snd_image_fst_preimage (f : X ⟶ Z) (g : Y ⟶ Z) (U : set X) : - (pullback.snd : pullback f g ⟶ _) '' ((pullback.fst : pullback f g ⟶ _) ⁻¹' U) = - g ⁻¹' (f '' U) := -begin - ext x, - split, - { rintros ⟨y, hy, rfl⟩, - exact ⟨(pullback.fst : pullback f g ⟶ _) y, hy, - concrete_category.congr_hom pullback.condition y⟩ }, - { rintros ⟨y, hy, eq⟩, - exact ⟨(Top.pullback_iso_prod_subtype f g).inv ⟨⟨_,_⟩, eq⟩, by simpa, by simp⟩ }, -end - -lemma pullback_fst_image_snd_preimage (f : X ⟶ Z) (g : Y ⟶ Z) (U : set Y) : - (pullback.fst : pullback f g ⟶ _) '' ((pullback.snd : pullback f g ⟶ _) ⁻¹' U) = - f ⁻¹' (g '' U) := -begin - ext x, - split, - { rintros ⟨y, hy, rfl⟩, - exact ⟨(pullback.snd : pullback f g ⟶ _) y, hy, - (concrete_category.congr_hom pullback.condition y).symm⟩ }, - { rintros ⟨y, hy, eq⟩, - exact ⟨(Top.pullback_iso_prod_subtype f g).inv ⟨⟨_,_⟩,eq.symm⟩, by simpa, by simp⟩ }, -end - -end pullback - -/-- The terminal object of `Top` is `punit`. -/ -def is_terminal_punit : is_terminal (Top.of punit.{u+1}) := -begin - haveI : ∀ X, unique (X ⟶ Top.of punit.{u+1}) := - λ X, ⟨⟨⟨λ x, punit.star, by continuity⟩⟩, λ f, by ext⟩, - exact limits.is_terminal.of_unique _, -end - -/-- The terminal object of `Top` is `punit`. -/ -def terminal_iso_punit : ⊤_ Top.{u} ≅ Top.of punit := -terminal_is_terminal.unique_up_to_iso is_terminal_punit - -/-- The initial object of `Top` is `pempty`. -/ -def is_initial_pempty : is_initial (Top.of pempty.{u+1}) := -begin - haveI : ∀ X, unique (Top.of pempty.{u+1} ⟶ X) := - λ X, ⟨⟨⟨λ x, x.elim, by continuity⟩⟩, λ f, by ext ⟨⟩⟩, - exact limits.is_initial.of_unique _, -end - -/-- The initial object of `Top` is `pempty`. -/ -def initial_iso_pempty : ⊥_ Top.{u} ≅ Top.of pempty := -initial_is_initial.unique_up_to_iso is_initial_pempty - -/-- The binary coproduct cofan in `Top`. -/ -protected -def binary_cofan (X Y : Top.{u}) : binary_cofan X Y := -binary_cofan.mk (⟨sum.inl⟩ : X ⟶ Top.of (X ⊕ Y)) ⟨sum.inr⟩ - -/-- The constructed binary coproduct cofan in `Top` is the coproduct. -/ -def binary_cofan_is_colimit (X Y : Top.{u}) : is_colimit (Top.binary_cofan X Y) := -begin - refine limits.binary_cofan.is_colimit_mk (λ s, ⟨sum.elim s.inl s.inr⟩) _ _ _, - { intro s, ext, refl }, - { intro s, ext, refl }, - { intros s m h₁ h₂, ext (x|x), - exacts [(concrete_category.congr_hom h₁ x : _), (concrete_category.congr_hom h₂ x : _)] }, -end - -lemma binary_cofan_is_colimit_iff {X Y : Top} (c : binary_cofan X Y) : - nonempty (is_colimit c) ↔ - open_embedding c.inl ∧ open_embedding c.inr ∧ is_compl (set.range c.inl) (set.range c.inr) := -begin - classical, - split, - { rintro ⟨h⟩, - rw [← show _ = c.inl, from h.comp_cocone_point_unique_up_to_iso_inv - (binary_cofan_is_colimit X Y) ⟨walking_pair.left⟩, - ← show _ = c.inr, from h.comp_cocone_point_unique_up_to_iso_inv - (binary_cofan_is_colimit X Y) ⟨walking_pair.right⟩], - dsimp, - refine - ⟨(homeo_of_iso $ h.cocone_point_unique_up_to_iso (binary_cofan_is_colimit X Y)).symm - .open_embedding.comp open_embedding_inl, (homeo_of_iso $ h.cocone_point_unique_up_to_iso - (binary_cofan_is_colimit X Y)).symm.open_embedding.comp open_embedding_inr, _⟩, - erw [set.range_comp, ← eq_compl_iff_is_compl, set.range_comp _ sum.inr, ← set.image_compl_eq - (homeo_of_iso $ h.cocone_point_unique_up_to_iso (binary_cofan_is_colimit X Y)) - .symm.bijective], - congr' 1, - exact set.compl_range_inr.symm }, - { rintros ⟨h₁, h₂, h₃⟩, - have : ∀ x, x ∈ set.range c.inl ∨ x ∈ set.range c.inr, - { rw [eq_compl_iff_is_compl.mpr h₃.symm], exact λ _, or_not }, - refine ⟨binary_cofan.is_colimit.mk _ _ _ _ _⟩, - { intros T f g, - refine continuous_map.mk _ _, - { exact λ x, if h : x ∈ set.range c.inl - then f ((equiv.of_injective _ h₁.inj).symm ⟨x, h⟩) - else g ((equiv.of_injective _ h₂.inj).symm ⟨x, (this x).resolve_left h⟩) }, - rw continuous_iff_continuous_at, - intro x, - by_cases x ∈ set.range c.inl, - { revert h x, - apply (is_open.continuous_on_iff _).mp, - { rw continuous_on_iff_continuous_restrict, - convert_to continuous (f ∘ (homeomorph.of_embedding _ h₁.to_embedding).symm), - { ext ⟨x, hx⟩, exact dif_pos hx }, - continuity }, - { exact h₁.open_range } }, - { revert h x, - apply (is_open.continuous_on_iff _).mp, - { rw continuous_on_iff_continuous_restrict, - have : ∀ a, a ∉ set.range c.inl → a ∈ set.range c.inr, - { rintros a (h : a ∈ (set.range c.inl)ᶜ), rwa eq_compl_iff_is_compl.mpr h₃.symm }, - convert_to continuous - (g ∘ (homeomorph.of_embedding _ h₂.to_embedding).symm ∘ subtype.map _ this), - { ext ⟨x, hx⟩, exact dif_neg hx }, - continuity, - rw embedding_subtype_coe.to_inducing.continuous_iff, - exact continuous_subtype_coe }, - { change is_open (set.range c.inl)ᶜ, rw ← eq_compl_iff_is_compl.mpr h₃.symm, - exact h₂.open_range } } }, - { intros T f g, ext x, refine (dif_pos _).trans _, { exact ⟨x, rfl⟩ }, - { rw equiv.of_injective_symm_apply } }, - { intros T f g, ext x, refine (dif_neg _).trans _, - { rintro ⟨y, e⟩, have : c.inr x ∈ set.range c.inl ⊓ set.range c.inr := ⟨⟨_, e⟩, ⟨_, rfl⟩⟩, - rwa disjoint_iff.mp h₃.1 at this }, - { exact congr_arg g (equiv.of_injective_symm_apply _ _) } }, - { rintro T _ _ m rfl rfl, ext x, change m x = dite _ _ _, - split_ifs; exact congr_arg _ (equiv.apply_of_injective_symm _ ⟨_, _⟩).symm } } -end - ---TODO: Add analogous constructions for `pushout`. - -lemma coinduced_of_is_colimit {F : J ⥤ Top.{max v u}} (c : cocone F) (hc : is_colimit c) : - c.X.topological_space = ⨆ j, (F.obj j).topological_space.coinduced (c.ι.app j) := -begin - let homeo := homeo_of_iso (hc.cocone_point_unique_up_to_iso (colimit_cocone_is_colimit F)), - ext, - refine homeo.symm.is_open_preimage.symm.trans (iff.trans _ is_open_supr_iff.symm), - exact is_open_supr_iff -end - -lemma colimit_topology (F : J ⥤ Top.{max v u}) : - (colimit F).topological_space = ⨆ j, (F.obj j).topological_space.coinduced (colimit.ι F j) := -coinduced_of_is_colimit _ (colimit.is_colimit F) - -lemma colimit_is_open_iff (F : J ⥤ Top.{max v u}) (U : set ((colimit F : _) : Type (max v u))) : - is_open U ↔ ∀ j, is_open (colimit.ι F j ⁻¹' U) := -begin - conv_lhs { rw colimit_topology F }, - exact is_open_supr_iff -end - -lemma coequalizer_is_open_iff (F : walking_parallel_pair ⥤ Top.{u}) - (U : set ((colimit F : _) : Type u)) : - is_open U ↔ is_open (colimit.ι F walking_parallel_pair.one ⁻¹' U) := -begin - rw colimit_is_open_iff.{u}, - split, - { intro H, exact H _ }, - { intros H j, - cases j, - { rw ←colimit.w F walking_parallel_pair_hom.left, - exact (F.map walking_parallel_pair_hom.left).continuous_to_fun.is_open_preimage _ H }, - { exact H } } -end - -end Top - -namespace Top - -section cofiltered_limit - -variables {J : Type v} [small_category J] [is_cofiltered J] (F : J ⥤ Top.{max v u}) - (C : cone F) (hC : is_limit C) - -include hC - -/-- -Given a *compatible* collection of topological bases for the factors in a cofiltered limit -which contain `set.univ` and are closed under intersections, the induced *naive* collection -of sets in the limit is, in fact, a topological basis. --/ -theorem is_topological_basis_cofiltered_limit - (T : Π j, set (set (F.obj j))) (hT : ∀ j, is_topological_basis (T j)) - (univ : ∀ (i : J), set.univ ∈ T i) - (inter : ∀ i (U1 U2 : set (F.obj i)), U1 ∈ T i → U2 ∈ T i → U1 ∩ U2 ∈ T i) - (compat : ∀ (i j : J) (f : i ⟶ j) (V : set (F.obj j)) (hV : V ∈ T j), (F.map f) ⁻¹' V ∈ T i) : - is_topological_basis { U : set C.X | ∃ j (V : set (F.obj j)), V ∈ T j ∧ U = C.π.app j ⁻¹' V } := -begin - classical, - -- The limit cone for `F` whose topology is defined as an infimum. - let D := limit_cone_infi F, - -- The isomorphism between the cone point of `C` and the cone point of `D`. - let E : C.X ≅ D.X := hC.cone_point_unique_up_to_iso (limit_cone_infi_is_limit _), - have hE : inducing E.hom := (Top.homeo_of_iso E).inducing, - -- Reduce to the assertion of the theorem with `D` instead of `C`. - suffices : is_topological_basis - { U : set D.X | ∃ j (V : set (F.obj j)), V ∈ T j ∧ U = D.π.app j ⁻¹' V }, - { convert this.inducing hE, - ext U0, - split, - { rintro ⟨j, V, hV, rfl⟩, - refine ⟨D.π.app j ⁻¹' V, ⟨j, V, hV, rfl⟩, rfl⟩ }, - { rintro ⟨W, ⟨j, V, hV, rfl⟩, rfl⟩, - refine ⟨j, V, hV, rfl⟩ } }, - -- Using `D`, we can apply the characterization of the topological basis of a - -- topology defined as an infimum... - convert is_topological_basis_infi hT (λ j (x : D.X), D.π.app j x), - ext U0, - split, - { rintros ⟨j, V, hV, rfl⟩, - let U : Π i, set (F.obj i) := λ i, if h : i = j then (by {rw h, exact V}) else set.univ, - refine ⟨U,{j},_,_⟩, - { rintro i h, - rw finset.mem_singleton at h, - dsimp [U], - rw dif_pos h, - subst h, - exact hV }, - { dsimp [U], - simp } }, - { rintros ⟨U, G, h1, h2⟩, - obtain ⟨j, hj⟩ := is_cofiltered.inf_objs_exists G, - let g : ∀ e (he : e ∈ G), j ⟶ e := λ _ he, (hj he).some, - let Vs : J → set (F.obj j) := λ e, if h : e ∈ G then F.map (g e h) ⁻¹' (U e) else set.univ, - let V : set (F.obj j) := ⋂ (e : J) (he : e ∈ G), Vs e, - refine ⟨j, V, _, _⟩, - { -- An intermediate claim used to apply induction along `G : finset J` later on. - have : ∀ (S : set (set (F.obj j))) (E : finset J) (P : J → set (F.obj j)) - (univ : set.univ ∈ S) - (inter : ∀ A B : set (F.obj j), A ∈ S → B ∈ S → A ∩ B ∈ S) - (cond : ∀ (e : J) (he : e ∈ E), P e ∈ S), (⋂ e (he : e ∈ E), P e) ∈ S, - { intros S E, - apply E.induction_on, - { intros P he hh, - simpa }, - { intros a E ha hh1 hh2 hh3 hh4 hh5, - rw finset.set_bInter_insert, - refine hh4 _ _ (hh5 _ (finset.mem_insert_self _ _)) (hh1 _ hh3 hh4 _), - intros e he, - exact hh5 e (finset.mem_insert_of_mem he) } }, - -- use the intermediate claim to finish off the goal using `univ` and `inter`. - refine this _ _ _ (univ _) (inter _) _, - intros e he, - dsimp [Vs], - rw dif_pos he, - exact compat j e (g e he) (U e) (h1 e he), }, - { -- conclude... - rw h2, - dsimp [V], - rw set.preimage_Inter, - congr' 1, - ext1 e, - rw set.preimage_Inter, - congr' 1, - ext1 he, - dsimp [Vs], - rw [dif_pos he, ← set.preimage_comp], - congr' 1, - change _ = ⇑(D.π.app j ≫ F.map (g e he)), - rw D.w } } -end - -end cofiltered_limit - -section topological_konig - -/-! -## Topological Kőnig's lemma - -A topological version of Kőnig's lemma is that the inverse limit of nonempty compact Hausdorff -spaces is nonempty. (Note: this can be generalized further to inverse limits of nonempty compact -T0 spaces, where all the maps are closed maps; see [Stone1979] --- however there is an erratum -for Theorem 4 that the element in the inverse limit can have cofinally many components that are -not closed points.) - -We give this in a more general form, which is that cofiltered limits -of nonempty compact Hausdorff spaces are nonempty -(`nonempty_limit_cone_of_compact_t2_cofiltered_system`). - -This also applies to inverse limits, where `{J : Type u} [preorder J] [is_directed J (≤)]` and -`F : Jᵒᵖ ⥤ Top`. - -The theorem is specialized to nonempty finite types (which are compact Hausdorff with the -discrete topology) in lemmas `nonempty_sections_of_finite_cofiltered_system` and -`nonempty_sections_of_finite_inverse_system` in the file `category_theory.cofiltered_system`. - -(See for the Set version.) --/ - -variables {J : Type u} [small_category J] -variables (F : J ⥤ Top.{u}) - -private abbreviation finite_diagram_arrow {J : Type u} [small_category J] (G : finset J) := -Σ' (X Y : J) (mX : X ∈ G) (mY : Y ∈ G), X ⟶ Y -private abbreviation finite_diagram (J : Type u) [small_category J] := -Σ (G : finset J), finset (finite_diagram_arrow G) - -/-- -Partial sections of a cofiltered limit are sections when restricted to -a finite subset of objects and morphisms of `J`. --/ -def partial_sections {J : Type u} [small_category J] (F : J ⥤ Top.{u}) - {G : finset J} (H : finset (finite_diagram_arrow G)) : set (Π j, F.obj j) := -{ u | ∀ {f : finite_diagram_arrow G} (hf : f ∈ H), F.map f.2.2.2.2 (u f.1) = u f.2.1 } - -lemma partial_sections.nonempty [is_cofiltered_or_empty J] [h : Π (j : J), nonempty (F.obj j)] - {G : finset J} (H : finset (finite_diagram_arrow G)) : - (partial_sections F H).nonempty := -begin - classical, - casesI is_empty_or_nonempty J, - { exact ⟨is_empty_elim, λ j, is_empty.elim' infer_instance j.1⟩ }, - haveI : is_cofiltered J := ⟨⟩, - use λ (j : J), if hj : j ∈ G - then F.map (is_cofiltered.inf_to G H hj) (h (is_cofiltered.inf G H)).some - else (h _).some, - rintros ⟨X, Y, hX, hY, f⟩ hf, - dsimp only, - rwa [dif_pos hX, dif_pos hY, ←comp_app, ←F.map_comp, - @is_cofiltered.inf_to_commutes _ _ _ G H], -end - -lemma partial_sections.directed : - directed superset (λ (G : finite_diagram J), partial_sections F G.2) := -begin - classical, - intros A B, - let ιA : finite_diagram_arrow A.1 → finite_diagram_arrow (A.1 ⊔ B.1) := - λ f, ⟨f.1, f.2.1, finset.mem_union_left _ f.2.2.1, finset.mem_union_left _ f.2.2.2.1, - f.2.2.2.2⟩, - let ιB : finite_diagram_arrow B.1 → finite_diagram_arrow (A.1 ⊔ B.1) := - λ f, ⟨f.1, f.2.1, finset.mem_union_right _ f.2.2.1, finset.mem_union_right _ f.2.2.2.1, - f.2.2.2.2⟩, - refine ⟨⟨A.1 ⊔ B.1, A.2.image ιA ⊔ B.2.image ιB⟩, _, _⟩, - { rintro u hu f hf, - have : ιA f ∈ A.2.image ιA ⊔ B.2.image ιB, - { apply finset.mem_union_left, - rw finset.mem_image, - refine ⟨f, hf, rfl⟩ }, - exact hu this }, - { rintro u hu f hf, - have : ιB f ∈ A.2.image ιA ⊔ B.2.image ιB, - { apply finset.mem_union_right, - rw finset.mem_image, - refine ⟨f, hf, rfl⟩ }, - exact hu this } -end - -lemma partial_sections.closed [Π (j : J), t2_space (F.obj j)] - {G : finset J} (H : finset (finite_diagram_arrow G)) : - is_closed (partial_sections F H) := -begin - have : partial_sections F H = - ⋂ {f : finite_diagram_arrow G} (hf : f ∈ H), { u | F.map f.2.2.2.2 (u f.1) = u f.2.1 }, - { ext1, - simp only [set.mem_Inter, set.mem_set_of_eq], - refl, }, - rw this, - apply is_closed_bInter, - intros f hf, - apply is_closed_eq, - continuity, -end - -/-- -Cofiltered limits of nonempty compact Hausdorff spaces are nonempty topological spaces. --/ -lemma nonempty_limit_cone_of_compact_t2_cofiltered_system - [is_cofiltered_or_empty J] - [Π (j : J), nonempty (F.obj j)] - [Π (j : J), compact_space (F.obj j)] - [Π (j : J), t2_space (F.obj j)] : - nonempty (Top.limit_cone.{u} F).X := -begin - classical, - obtain ⟨u, hu⟩ := is_compact.nonempty_Inter_of_directed_nonempty_compact_closed - (λ G, partial_sections F _) - (partial_sections.directed F) - (λ G, partial_sections.nonempty F _) - (λ G, is_closed.is_compact (partial_sections.closed F _)) - (λ G, partial_sections.closed F _), - use u, - intros X Y f, - let G : finite_diagram J := - ⟨{X, Y}, - {⟨X, Y, - by simp only [true_or, eq_self_iff_true, finset.mem_insert], - by simp only [eq_self_iff_true, or_true, finset.mem_insert, finset.mem_singleton], - f⟩}⟩, - exact hu _ ⟨G, rfl⟩ (finset.mem_singleton_self _), -end - -end topological_konig - -end Top diff --git a/src/topology/category/Top/limits/basic.lean b/src/topology/category/Top/limits/basic.lean new file mode 100644 index 0000000000000..68650032113db --- /dev/null +++ b/src/topology/category/Top/limits/basic.lean @@ -0,0 +1,162 @@ +/- +Copyright (c) 2017 Scott Morrison. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Patrick Massot, Scott Morrison, Mario Carneiro, Andrew Yang +-/ +import topology.category.Top.basic +import category_theory.limits.concrete_category + +/-! +# The category of topological spaces has all limits and colimits + +Further, these limits and colimits are preserved by the forgetful functor --- that is, the +underlying types are just the limits in the category of types. +-/ + +open topological_space +open category_theory +open category_theory.limits +open opposite + +universes u v w + +noncomputable theory + +namespace Top + +variables {J : Type v} [small_category J] + +local notation `forget` := forget Top + +/-- +A choice of limit cone for a functor `F : J ⥤ Top`. +Generally you should just use `limit.cone F`, unless you need the actual definition +(which is in terms of `types.limit_cone`). +-/ +def limit_cone (F : J ⥤ Top.{max v u}) : cone F := +{ X := Top.of {u : Π j : J, F.obj j | ∀ {i j : J} (f : i ⟶ j), F.map f (u i) = u j}, + π := + { app := λ j, + { to_fun := λ u, u.val j, + continuous_to_fun := show continuous ((λ u : Π j : J, F.obj j, u j) ∘ subtype.val), + by continuity } } } + +/-- +A choice of limit cone for a functor `F : J ⥤ Top` whose topology is defined as an +infimum of topologies infimum. +Generally you should just use `limit.cone F`, unless you need the actual definition +(which is in terms of `types.limit_cone`). +-/ +def limit_cone_infi (F : J ⥤ Top.{max v u}) : cone F := +{ X := ⟨(types.limit_cone (F ⋙ forget)).X, ⨅j, + (F.obj j).str.induced ((types.limit_cone (F ⋙ forget)).π.app j)⟩, + π := + { app := λ j, ⟨(types.limit_cone (F ⋙ forget)).π.app j, + continuous_iff_le_induced.mpr (infi_le _ _)⟩, + naturality' := λ j j' f, continuous_map.coe_injective + ((types.limit_cone (F ⋙ forget)).π.naturality f) } } + +/-- +The chosen cone `Top.limit_cone F` for a functor `F : J ⥤ Top` is a limit cone. +Generally you should just use `limit.is_limit F`, unless you need the actual definition +(which is in terms of `types.limit_cone_is_limit`). +-/ +def limit_cone_is_limit (F : J ⥤ Top.{max v u}) : is_limit (limit_cone F) := +{ lift := λ S, { to_fun := λ x, ⟨λ j, S.π.app _ x, λ i j f, by { dsimp, erw ← S.w f, refl }⟩ }, + uniq' := λ S m h, by { ext : 3, simpa [← h] } } + +/-- +The chosen cone `Top.limit_cone_infi F` for a functor `F : J ⥤ Top` is a limit cone. +Generally you should just use `limit.is_limit F`, unless you need the actual definition +(which is in terms of `types.limit_cone_is_limit`). +-/ +def limit_cone_infi_is_limit (F : J ⥤ Top.{max v u}) : is_limit (limit_cone_infi F) := +by { refine is_limit.of_faithful forget (types.limit_cone_is_limit _) (λ s, ⟨_, _⟩) (λ s, rfl), + exact continuous_iff_coinduced_le.mpr (le_infi $ λ j, + coinduced_le_iff_le_induced.mp $ (continuous_iff_coinduced_le.mp (s.π.app j).continuous : + _) ) } + +instance Top_has_limits_of_size : has_limits_of_size.{v} Top.{max v u} := +{ has_limits_of_shape := λ J 𝒥, by exactI + { has_limit := λ F, has_limit.mk { cone := limit_cone F, is_limit := limit_cone_is_limit F } } } + +instance Top_has_limits : has_limits Top.{u} := Top.Top_has_limits_of_size.{u u} + +instance forget_preserves_limits_of_size : + preserves_limits_of_size.{v v} (forget : Top.{max v u} ⥤ Type (max v u)) := +{ preserves_limits_of_shape := λ J 𝒥, + { preserves_limit := λ F, + by exactI preserves_limit_of_preserves_limit_cone + (limit_cone_is_limit F) (types.limit_cone_is_limit (F ⋙ forget)) } } + +instance forget_preserves_limits : preserves_limits (forget : Top.{u} ⥤ Type u) := +Top.forget_preserves_limits_of_size.{u u} + +/-- +A choice of colimit cocone for a functor `F : J ⥤ Top`. +Generally you should just use `colimit.coone F`, unless you need the actual definition +(which is in terms of `types.colimit_cocone`). +-/ +def colimit_cocone (F : J ⥤ Top.{max v u}) : cocone F := +{ X := ⟨(types.colimit_cocone (F ⋙ forget)).X, ⨆ j, + (F.obj j).str.coinduced ((types.colimit_cocone (F ⋙ forget)).ι.app j)⟩, + ι := + { app := λ j, ⟨(types.colimit_cocone (F ⋙ forget)).ι.app j, + continuous_iff_coinduced_le.mpr (le_supr _ j)⟩, + naturality' := λ j j' f, continuous_map.coe_injective + ((types.colimit_cocone (F ⋙ forget)).ι.naturality f) } } + +/-- +The chosen cocone `Top.colimit_cocone F` for a functor `F : J ⥤ Top` is a colimit cocone. +Generally you should just use `colimit.is_colimit F`, unless you need the actual definition +(which is in terms of `types.colimit_cocone_is_colimit`). +-/ +def colimit_cocone_is_colimit (F : J ⥤ Top.{max v u}) : is_colimit (colimit_cocone F) := +by { refine is_colimit.of_faithful forget (types.colimit_cocone_is_colimit _) (λ s, ⟨_, _⟩) + (λ s, rfl), + exact continuous_iff_le_induced.mpr (supr_le $ λ j, + coinduced_le_iff_le_induced.mp $ (continuous_iff_coinduced_le.mp (s.ι.app j).continuous : + _) ) } + +instance Top_has_colimits_of_size : has_colimits_of_size.{v} Top.{max v u} := +{ has_colimits_of_shape := λ J 𝒥, by exactI + { has_colimit := λ F, has_colimit.mk { cocone := colimit_cocone F, is_colimit := + colimit_cocone_is_colimit F } } } + +instance Top_has_colimits : has_colimits Top.{u} := Top.Top_has_colimits_of_size.{u u} + +instance forget_preserves_colimits_of_size : + preserves_colimits_of_size.{v v} (forget : Top.{max v u} ⥤ Type (max v u)) := +{ preserves_colimits_of_shape := λ J 𝒥, + { preserves_colimit := λ F, + by exactI preserves_colimit_of_preserves_colimit_cocone + (colimit_cocone_is_colimit F) (types.colimit_cocone_is_colimit (F ⋙ forget)) } } + +instance forget_preserves_colimits : preserves_colimits (forget : Top.{u} ⥤ Type u) := +Top.forget_preserves_colimits_of_size.{u u} + +/-- The terminal object of `Top` is `punit`. -/ +def is_terminal_punit : is_terminal (Top.of punit.{u+1}) := +begin + haveI : ∀ X, unique (X ⟶ Top.of punit.{u+1}) := + λ X, ⟨⟨⟨λ x, punit.star, by continuity⟩⟩, λ f, by ext⟩, + exact limits.is_terminal.of_unique _, +end + +/-- The terminal object of `Top` is `punit`. -/ +def terminal_iso_punit : ⊤_ Top.{u} ≅ Top.of punit := +terminal_is_terminal.unique_up_to_iso is_terminal_punit + +/-- The initial object of `Top` is `pempty`. -/ +def is_initial_pempty : is_initial (Top.of pempty.{u+1}) := +begin + haveI : ∀ X, unique (Top.of pempty.{u+1} ⟶ X) := + λ X, ⟨⟨⟨λ x, x.elim, by continuity⟩⟩, λ f, by ext ⟨⟩⟩, + exact limits.is_initial.of_unique _, +end + +/-- The initial object of `Top` is `pempty`. -/ +def initial_iso_pempty : ⊥_ Top.{u} ≅ Top.of pempty := +initial_is_initial.unique_up_to_iso is_initial_pempty + +end Top diff --git a/src/topology/category/Top/limits/cofiltered.lean b/src/topology/category/Top/limits/cofiltered.lean new file mode 100644 index 0000000000000..980f9f23f4fe4 --- /dev/null +++ b/src/topology/category/Top/limits/cofiltered.lean @@ -0,0 +1,121 @@ +/- +Copyright (c) 2017 Scott Morrison. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Patrick Massot, Scott Morrison, Mario Carneiro, Andrew Yang +-/ +import topology.category.Top.limits.basic + +/-! +# Cofiltered limits in Top. + +Given a *compatible* collection of topological bases for the factors in a cofiltered limit +which contain `set.univ` and are closed under intersections, the induced *naive* collection +of sets in the limit is, in fact, a topological basis. +-/ + +open topological_space +open category_theory +open category_theory.limits + +universes u v w + +noncomputable theory + +namespace Top + +section cofiltered_limit + +variables {J : Type v} [small_category J] [is_cofiltered J] (F : J ⥤ Top.{max v u}) + (C : cone F) (hC : is_limit C) + +include hC + +/-- +Given a *compatible* collection of topological bases for the factors in a cofiltered limit +which contain `set.univ` and are closed under intersections, the induced *naive* collection +of sets in the limit is, in fact, a topological basis. +-/ +theorem is_topological_basis_cofiltered_limit + (T : Π j, set (set (F.obj j))) (hT : ∀ j, is_topological_basis (T j)) + (univ : ∀ (i : J), set.univ ∈ T i) + (inter : ∀ i (U1 U2 : set (F.obj i)), U1 ∈ T i → U2 ∈ T i → U1 ∩ U2 ∈ T i) + (compat : ∀ (i j : J) (f : i ⟶ j) (V : set (F.obj j)) (hV : V ∈ T j), (F.map f) ⁻¹' V ∈ T i) : + is_topological_basis { U : set C.X | ∃ j (V : set (F.obj j)), V ∈ T j ∧ U = C.π.app j ⁻¹' V } := +begin + classical, + -- The limit cone for `F` whose topology is defined as an infimum. + let D := limit_cone_infi F, + -- The isomorphism between the cone point of `C` and the cone point of `D`. + let E : C.X ≅ D.X := hC.cone_point_unique_up_to_iso (limit_cone_infi_is_limit _), + have hE : inducing E.hom := (Top.homeo_of_iso E).inducing, + -- Reduce to the assertion of the theorem with `D` instead of `C`. + suffices : is_topological_basis + { U : set D.X | ∃ j (V : set (F.obj j)), V ∈ T j ∧ U = D.π.app j ⁻¹' V }, + { convert this.inducing hE, + ext U0, + split, + { rintro ⟨j, V, hV, rfl⟩, + refine ⟨D.π.app j ⁻¹' V, ⟨j, V, hV, rfl⟩, rfl⟩ }, + { rintro ⟨W, ⟨j, V, hV, rfl⟩, rfl⟩, + refine ⟨j, V, hV, rfl⟩ } }, + -- Using `D`, we can apply the characterization of the topological basis of a + -- topology defined as an infimum... + convert is_topological_basis_infi hT (λ j (x : D.X), D.π.app j x), + ext U0, + split, + { rintros ⟨j, V, hV, rfl⟩, + let U : Π i, set (F.obj i) := λ i, if h : i = j then (by {rw h, exact V}) else set.univ, + refine ⟨U,{j},_,_⟩, + { rintro i h, + rw finset.mem_singleton at h, + dsimp [U], + rw dif_pos h, + subst h, + exact hV }, + { dsimp [U], + simp } }, + { rintros ⟨U, G, h1, h2⟩, + obtain ⟨j, hj⟩ := is_cofiltered.inf_objs_exists G, + let g : ∀ e (he : e ∈ G), j ⟶ e := λ _ he, (hj he).some, + let Vs : J → set (F.obj j) := λ e, if h : e ∈ G then F.map (g e h) ⁻¹' (U e) else set.univ, + let V : set (F.obj j) := ⋂ (e : J) (he : e ∈ G), Vs e, + refine ⟨j, V, _, _⟩, + { -- An intermediate claim used to apply induction along `G : finset J` later on. + have : ∀ (S : set (set (F.obj j))) (E : finset J) (P : J → set (F.obj j)) + (univ : set.univ ∈ S) + (inter : ∀ A B : set (F.obj j), A ∈ S → B ∈ S → A ∩ B ∈ S) + (cond : ∀ (e : J) (he : e ∈ E), P e ∈ S), (⋂ e (he : e ∈ E), P e) ∈ S, + { intros S E, + apply E.induction_on, + { intros P he hh, + simpa }, + { intros a E ha hh1 hh2 hh3 hh4 hh5, + rw finset.set_bInter_insert, + refine hh4 _ _ (hh5 _ (finset.mem_insert_self _ _)) (hh1 _ hh3 hh4 _), + intros e he, + exact hh5 e (finset.mem_insert_of_mem he) } }, + -- use the intermediate claim to finish off the goal using `univ` and `inter`. + refine this _ _ _ (univ _) (inter _) _, + intros e he, + dsimp [Vs], + rw dif_pos he, + exact compat j e (g e he) (U e) (h1 e he), }, + { -- conclude... + rw h2, + dsimp [V], + rw set.preimage_Inter, + congr' 1, + ext1 e, + rw set.preimage_Inter, + congr' 1, + ext1 he, + dsimp [Vs], + rw [dif_pos he, ← set.preimage_comp], + congr' 1, + change _ = ⇑(D.π.app j ≫ F.map (g e he)), + rw D.w } } +end + +end cofiltered_limit + +end Top diff --git a/src/topology/category/Top/limits/konig.lean b/src/topology/category/Top/limits/konig.lean new file mode 100644 index 0000000000000..0548dc673a9b2 --- /dev/null +++ b/src/topology/category/Top/limits/konig.lean @@ -0,0 +1,147 @@ +/- +Copyright (c) 2017 Scott Morrison. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Patrick Massot, Scott Morrison, Mario Carneiro, Andrew Yang +-/ +import topology.category.Top.limits.basic + +/-! +# Topological Kőnig's lemma + +A topological version of Kőnig's lemma is that the inverse limit of nonempty compact Hausdorff +spaces is nonempty. (Note: this can be generalized further to inverse limits of nonempty compact +T0 spaces, where all the maps are closed maps; see [Stone1979] --- however there is an erratum +for Theorem 4 that the element in the inverse limit can have cofinally many components that are +not closed points.) + +We give this in a more general form, which is that cofiltered limits +of nonempty compact Hausdorff spaces are nonempty +(`nonempty_limit_cone_of_compact_t2_cofiltered_system`). + +This also applies to inverse limits, where `{J : Type u} [preorder J] [is_directed J (≤)]` and +`F : Jᵒᵖ ⥤ Top`. + +The theorem is specialized to nonempty finite types (which are compact Hausdorff with the +discrete topology) in lemmas `nonempty_sections_of_finite_cofiltered_system` and +`nonempty_sections_of_finite_inverse_system` in the file `category_theory.cofiltered_system`. + +(See for the Set version.) +-/ + +open category_theory +open category_theory.limits + +universes u v w + +noncomputable theory + +namespace Top + +section topological_konig + +variables {J : Type u} [small_category J] +variables (F : J ⥤ Top.{u}) + +private abbreviation finite_diagram_arrow {J : Type u} [small_category J] (G : finset J) := +Σ' (X Y : J) (mX : X ∈ G) (mY : Y ∈ G), X ⟶ Y +private abbreviation finite_diagram (J : Type u) [small_category J] := +Σ (G : finset J), finset (finite_diagram_arrow G) + +/-- +Partial sections of a cofiltered limit are sections when restricted to +a finite subset of objects and morphisms of `J`. +-/ +def partial_sections {J : Type u} [small_category J] (F : J ⥤ Top.{u}) + {G : finset J} (H : finset (finite_diagram_arrow G)) : set (Π j, F.obj j) := +{ u | ∀ {f : finite_diagram_arrow G} (hf : f ∈ H), F.map f.2.2.2.2 (u f.1) = u f.2.1 } + +lemma partial_sections.nonempty [is_cofiltered_or_empty J] [h : Π (j : J), nonempty (F.obj j)] + {G : finset J} (H : finset (finite_diagram_arrow G)) : + (partial_sections F H).nonempty := +begin + classical, + casesI is_empty_or_nonempty J, + { exact ⟨is_empty_elim, λ j, is_empty.elim' infer_instance j.1⟩ }, + haveI : is_cofiltered J := ⟨⟩, + use λ (j : J), if hj : j ∈ G + then F.map (is_cofiltered.inf_to G H hj) (h (is_cofiltered.inf G H)).some + else (h _).some, + rintros ⟨X, Y, hX, hY, f⟩ hf, + dsimp only, + rwa [dif_pos hX, dif_pos hY, ←comp_app, ←F.map_comp, + @is_cofiltered.inf_to_commutes _ _ _ G H], +end + +lemma partial_sections.directed : + directed superset (λ (G : finite_diagram J), partial_sections F G.2) := +begin + classical, + intros A B, + let ιA : finite_diagram_arrow A.1 → finite_diagram_arrow (A.1 ⊔ B.1) := + λ f, ⟨f.1, f.2.1, finset.mem_union_left _ f.2.2.1, finset.mem_union_left _ f.2.2.2.1, + f.2.2.2.2⟩, + let ιB : finite_diagram_arrow B.1 → finite_diagram_arrow (A.1 ⊔ B.1) := + λ f, ⟨f.1, f.2.1, finset.mem_union_right _ f.2.2.1, finset.mem_union_right _ f.2.2.2.1, + f.2.2.2.2⟩, + refine ⟨⟨A.1 ⊔ B.1, A.2.image ιA ⊔ B.2.image ιB⟩, _, _⟩, + { rintro u hu f hf, + have : ιA f ∈ A.2.image ιA ⊔ B.2.image ιB, + { apply finset.mem_union_left, + rw finset.mem_image, + refine ⟨f, hf, rfl⟩ }, + exact hu this }, + { rintro u hu f hf, + have : ιB f ∈ A.2.image ιA ⊔ B.2.image ιB, + { apply finset.mem_union_right, + rw finset.mem_image, + refine ⟨f, hf, rfl⟩ }, + exact hu this } +end + +lemma partial_sections.closed [Π (j : J), t2_space (F.obj j)] + {G : finset J} (H : finset (finite_diagram_arrow G)) : + is_closed (partial_sections F H) := +begin + have : partial_sections F H = + ⋂ {f : finite_diagram_arrow G} (hf : f ∈ H), { u | F.map f.2.2.2.2 (u f.1) = u f.2.1 }, + { ext1, + simp only [set.mem_Inter, set.mem_set_of_eq], + refl, }, + rw this, + apply is_closed_bInter, + intros f hf, + apply is_closed_eq, + continuity, +end + +/-- +Cofiltered limits of nonempty compact Hausdorff spaces are nonempty topological spaces. +-/ +lemma nonempty_limit_cone_of_compact_t2_cofiltered_system + [is_cofiltered_or_empty J] + [Π (j : J), nonempty (F.obj j)] + [Π (j : J), compact_space (F.obj j)] + [Π (j : J), t2_space (F.obj j)] : + nonempty (Top.limit_cone.{u} F).X := +begin + classical, + obtain ⟨u, hu⟩ := is_compact.nonempty_Inter_of_directed_nonempty_compact_closed + (λ G, partial_sections F _) + (partial_sections.directed F) + (λ G, partial_sections.nonempty F _) + (λ G, is_closed.is_compact (partial_sections.closed F _)) + (λ G, partial_sections.closed F _), + use u, + intros X Y f, + let G : finite_diagram J := + ⟨{X, Y}, + {⟨X, Y, + by simp only [true_or, eq_self_iff_true, finset.mem_insert], + by simp only [eq_self_iff_true, or_true, finset.mem_insert, finset.mem_singleton], + f⟩}⟩, + exact hu _ ⟨G, rfl⟩ (finset.mem_singleton_self _), +end + +end topological_konig + +end Top diff --git a/src/topology/category/Top/limits/products.lean b/src/topology/category/Top/limits/products.lean new file mode 100644 index 0000000000000..fa5c7931f9f69 --- /dev/null +++ b/src/topology/category/Top/limits/products.lean @@ -0,0 +1,311 @@ +/- +Copyright (c) 2017 Scott Morrison. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Patrick Massot, Scott Morrison, Mario Carneiro, Andrew Yang +-/ +import topology.category.Top.epi_mono +import topology.category.Top.limits.basic + +/-! +# Products and coproducts in the category of topological spaces + +-/ + +open topological_space +open category_theory +open category_theory.limits + +universes u v w + +noncomputable theory + +namespace Top + +variables {J : Type v} [small_category J] + +/-- The projection from the product as a bundled continous map. -/ +abbreviation pi_π {ι : Type v} (α : ι → Top.{max v u}) (i : ι) : Top.of (Π i, α i) ⟶ α i := +⟨λ f, f i, continuous_apply i⟩ + +/-- The explicit fan of a family of topological spaces given by the pi type. -/ +@[simps X π_app] +def pi_fan {ι : Type v} (α : ι → Top.{max v u}) : fan α := +fan.mk (Top.of (Π i, α i)) (pi_π α) + +/-- The constructed fan is indeed a limit -/ +def pi_fan_is_limit {ι : Type v} (α : ι → Top.{max v u}) : is_limit (pi_fan α) := +{ lift := λ S, { to_fun := λ s i, S.π.app ⟨i⟩ s }, + uniq' := by { intros S m h, ext x i, simp [← h ⟨i⟩] }, + fac' := λ s j, by { cases j, tidy, }, } + +/-- +The product is homeomorphic to the product of the underlying spaces, +equipped with the product topology. +-/ +def pi_iso_pi {ι : Type v} (α : ι → Top.{max v u}) : ∏ α ≅ Top.of (Π i, α i) := +(limit.is_limit _).cone_point_unique_up_to_iso (pi_fan_is_limit α) + +@[simp, reassoc] +lemma pi_iso_pi_inv_π {ι : Type v} (α : ι → Top.{max v u}) (i : ι) : + (pi_iso_pi α).inv ≫ pi.π α i = pi_π α i := +by simp [pi_iso_pi] + +@[simp] +lemma pi_iso_pi_inv_π_apply {ι : Type v} (α : ι → Top.{max v u}) (i : ι) (x : Π i, α i) : + (pi.π α i : _) ((pi_iso_pi α).inv x) = x i := +concrete_category.congr_hom (pi_iso_pi_inv_π α i) x + +@[simp] +lemma pi_iso_pi_hom_apply {ι : Type v} (α : ι → Top.{max v u}) (i : ι) (x : ∏ α) : + (pi_iso_pi α).hom x i = (pi.π α i : _) x := +begin + have := pi_iso_pi_inv_π α i, + rw iso.inv_comp_eq at this, + exact concrete_category.congr_hom this x +end + +/-- The inclusion to the coproduct as a bundled continous map. -/ +abbreviation sigma_ι {ι : Type v} (α : ι → Top.{max v u}) (i : ι) : α i ⟶ Top.of (Σ i, α i) := +⟨sigma.mk i⟩ + +/-- The explicit cofan of a family of topological spaces given by the sigma type. -/ +@[simps X ι_app] +def sigma_cofan {ι : Type v} (α : ι → Top.{max v u}) : cofan α := +cofan.mk (Top.of (Σ i, α i)) (sigma_ι α) + +/-- The constructed cofan is indeed a colimit -/ +def sigma_cofan_is_colimit {ι : Type v} (α : ι → Top.{max v u}) : is_colimit (sigma_cofan α) := +{ desc := λ S, { to_fun := λ s, S.ι.app ⟨s.1⟩ s.2, + continuous_to_fun := continuous_sigma $ λ i, map_continuous (S.ι.app ⟨i⟩) }, + uniq' := by { intros S m h, ext ⟨i, x⟩, simp [← h ⟨i⟩] }, + fac' := λ s j, by { cases j, tidy, }, } + +/-- +The coproduct is homeomorphic to the disjoint union of the topological spaces. +-/ +def sigma_iso_sigma {ι : Type v} (α : ι → Top.{max v u}) : ∐ α ≅ Top.of (Σ i, α i) := +(colimit.is_colimit _).cocone_point_unique_up_to_iso (sigma_cofan_is_colimit α) + +@[simp, reassoc] +lemma sigma_iso_sigma_hom_ι {ι : Type v} (α : ι → Top.{max v u}) (i : ι) : + sigma.ι α i ≫ (sigma_iso_sigma α).hom = sigma_ι α i := +by simp [sigma_iso_sigma] + +@[simp] +lemma sigma_iso_sigma_hom_ι_apply {ι : Type v} (α : ι → Top.{max v u}) (i : ι) (x : α i) : + (sigma_iso_sigma α).hom ((sigma.ι α i : _) x) = sigma.mk i x := +concrete_category.congr_hom (sigma_iso_sigma_hom_ι α i) x + +@[simp] +lemma sigma_iso_sigma_inv_apply {ι : Type v} (α : ι → Top.{max v u}) (i : ι) (x : α i) : + (sigma_iso_sigma α).inv ⟨i, x⟩ = (sigma.ι α i : _) x := +by { rw [← sigma_iso_sigma_hom_ι_apply, ← comp_app], simp, } + +lemma induced_of_is_limit {F : J ⥤ Top.{max v u}} (C : cone F) (hC : is_limit C) : + C.X.topological_space = ⨅ j, (F.obj j).topological_space.induced (C.π.app j) := +begin + let homeo := homeo_of_iso (hC.cone_point_unique_up_to_iso (limit_cone_infi_is_limit F)), + refine homeo.inducing.induced.trans _, + change induced homeo (⨅ (j : J), _) = _, + simpa [induced_infi, induced_compose], +end + +lemma limit_topology (F : J ⥤ Top.{max v u}) : + (limit F).topological_space = ⨅ j, (F.obj j).topological_space.induced (limit.π F j) := +induced_of_is_limit _ (limit.is_limit F) + +section prod + +/-- The first projection from the product. -/ +abbreviation prod_fst {X Y : Top.{u}} : Top.of (X × Y) ⟶ X := ⟨prod.fst⟩ + +/-- The second projection from the product. -/ +abbreviation prod_snd {X Y : Top.{u}} : Top.of (X × Y) ⟶ Y := ⟨prod.snd⟩ + +/-- The explicit binary cofan of `X, Y` given by `X × Y`. -/ +def prod_binary_fan (X Y : Top.{u}) : binary_fan X Y := +binary_fan.mk prod_fst prod_snd + +/-- The constructed binary fan is indeed a limit -/ +def prod_binary_fan_is_limit (X Y : Top.{u}) : is_limit (prod_binary_fan X Y) := +{ lift := λ (S : binary_fan X Y), { to_fun := λ s, (S.fst s, S.snd s) }, + fac' := begin + rintros S (_|_), + tidy + end, + uniq' := begin + intros S m h, + ext x, + { specialize h ⟨walking_pair.left⟩, + apply_fun (λ e, (e x)) at h, + exact h }, + { specialize h ⟨walking_pair.right⟩, + apply_fun (λ e, (e x)) at h, + exact h }, + end } + +/-- +The homeomorphism between `X ⨯ Y` and the set-theoretic product of `X` and `Y`, +equipped with the product topology. +-/ +def prod_iso_prod (X Y : Top.{u}) : X ⨯ Y ≅ Top.of (X × Y) := +(limit.is_limit _).cone_point_unique_up_to_iso (prod_binary_fan_is_limit X Y) + +@[simp, reassoc] lemma prod_iso_prod_hom_fst (X Y : Top.{u}) : + (prod_iso_prod X Y).hom ≫ prod_fst = limits.prod.fst := +by simpa [← iso.eq_inv_comp, prod_iso_prod] + +@[simp, reassoc] lemma prod_iso_prod_hom_snd (X Y : Top.{u}) : + (prod_iso_prod X Y).hom ≫ prod_snd = limits.prod.snd := +by simpa [← iso.eq_inv_comp, prod_iso_prod] + +@[simp] lemma prod_iso_prod_hom_apply {X Y : Top.{u}} (x : X ⨯ Y) : + (prod_iso_prod X Y).hom x = + ((limits.prod.fst : X ⨯ Y ⟶ _) x, (limits.prod.snd : X ⨯ Y ⟶ _) x) := +begin + ext, + { exact concrete_category.congr_hom (prod_iso_prod_hom_fst X Y) x }, + { exact concrete_category.congr_hom (prod_iso_prod_hom_snd X Y) x } +end + +@[simp, reassoc, elementwise] lemma prod_iso_prod_inv_fst (X Y : Top.{u}) : + (prod_iso_prod X Y).inv ≫ limits.prod.fst = prod_fst := +by simp [iso.inv_comp_eq] + +@[simp, reassoc, elementwise] lemma prod_iso_prod_inv_snd (X Y : Top.{u}) : + (prod_iso_prod X Y).inv ≫ limits.prod.snd = prod_snd := +by simp [iso.inv_comp_eq] + +lemma prod_topology {X Y : Top} : + (X ⨯ Y).topological_space = + induced (limits.prod.fst : X ⨯ Y ⟶ _) X.topological_space ⊓ + induced (limits.prod.snd : X ⨯ Y ⟶ _) Y.topological_space := +begin + let homeo := homeo_of_iso (prod_iso_prod X Y), + refine homeo.inducing.induced.trans _, + change induced homeo (_ ⊓ _) = _, + simpa [induced_compose] +end + +lemma range_prod_map {W X Y Z : Top.{u}} (f : W ⟶ Y) (g : X ⟶ Z) : + set.range (limits.prod.map f g) = + (limits.prod.fst : Y ⨯ Z ⟶ _) ⁻¹' (set.range f) ∩ + (limits.prod.snd : Y ⨯ Z ⟶ _) ⁻¹' (set.range g) := +begin + ext, + split, + { rintros ⟨y, rfl⟩, + simp only [set.mem_preimage, set.mem_range, set.mem_inter_iff, ←comp_apply], + simp only [limits.prod.map_fst, limits.prod.map_snd, + exists_apply_eq_apply, comp_apply, and_self] }, + { rintros ⟨⟨x₁, hx₁⟩, ⟨x₂, hx₂⟩⟩, + use (prod_iso_prod W X).inv (x₁, x₂), + apply concrete.limit_ext, + rintro ⟨⟨⟩⟩, + { simp only [← comp_apply, category.assoc], erw limits.prod.map_fst, simp [hx₁] }, + { simp only [← comp_apply, category.assoc], erw limits.prod.map_snd, simp [hx₂] } } +end + +lemma inducing_prod_map {W X Y Z : Top} {f : W ⟶ X} {g : Y ⟶ Z} + (hf : inducing f) (hg : inducing g) : inducing (limits.prod.map f g) := +begin + constructor, + simp only [prod_topology, induced_compose, ←coe_comp, limits.prod.map_fst, limits.prod.map_snd, + induced_inf], + simp only [coe_comp], + rw [← @induced_compose _ _ _ _ _ f, ← @induced_compose _ _ _ _ _ g, ← hf.induced, ← hg.induced] +end + +lemma embedding_prod_map {W X Y Z : Top} {f : W ⟶ X} {g : Y ⟶ Z} + (hf : embedding f) (hg : embedding g) : embedding (limits.prod.map f g) := +⟨inducing_prod_map hf.to_inducing hg.to_inducing, +begin + haveI := (Top.mono_iff_injective _).mpr hf.inj, + haveI := (Top.mono_iff_injective _).mpr hg.inj, + exact (Top.mono_iff_injective _).mp infer_instance +end⟩ + +end prod + + +/-- The binary coproduct cofan in `Top`. -/ +protected +def binary_cofan (X Y : Top.{u}) : binary_cofan X Y := +binary_cofan.mk (⟨sum.inl⟩ : X ⟶ Top.of (X ⊕ Y)) ⟨sum.inr⟩ + +/-- The constructed binary coproduct cofan in `Top` is the coproduct. -/ +def binary_cofan_is_colimit (X Y : Top.{u}) : is_colimit (Top.binary_cofan X Y) := +begin + refine limits.binary_cofan.is_colimit_mk (λ s, ⟨sum.elim s.inl s.inr⟩) _ _ _, + { intro s, ext, refl }, + { intro s, ext, refl }, + { intros s m h₁ h₂, ext (x|x), + exacts [(concrete_category.congr_hom h₁ x : _), (concrete_category.congr_hom h₂ x : _)] }, +end + +lemma binary_cofan_is_colimit_iff {X Y : Top} (c : binary_cofan X Y) : + nonempty (is_colimit c) ↔ + open_embedding c.inl ∧ open_embedding c.inr ∧ is_compl (set.range c.inl) (set.range c.inr) := +begin + classical, + split, + { rintro ⟨h⟩, + rw [← show _ = c.inl, from h.comp_cocone_point_unique_up_to_iso_inv + (binary_cofan_is_colimit X Y) ⟨walking_pair.left⟩, + ← show _ = c.inr, from h.comp_cocone_point_unique_up_to_iso_inv + (binary_cofan_is_colimit X Y) ⟨walking_pair.right⟩], + dsimp, + refine + ⟨(homeo_of_iso $ h.cocone_point_unique_up_to_iso (binary_cofan_is_colimit X Y)).symm + .open_embedding.comp open_embedding_inl, (homeo_of_iso $ h.cocone_point_unique_up_to_iso + (binary_cofan_is_colimit X Y)).symm.open_embedding.comp open_embedding_inr, _⟩, + erw [set.range_comp, ← eq_compl_iff_is_compl, set.range_comp _ sum.inr, ← set.image_compl_eq + (homeo_of_iso $ h.cocone_point_unique_up_to_iso (binary_cofan_is_colimit X Y)) + .symm.bijective], + congr' 1, + exact set.compl_range_inr.symm }, + { rintros ⟨h₁, h₂, h₃⟩, + have : ∀ x, x ∈ set.range c.inl ∨ x ∈ set.range c.inr, + { rw [eq_compl_iff_is_compl.mpr h₃.symm], exact λ _, or_not }, + refine ⟨binary_cofan.is_colimit.mk _ _ _ _ _⟩, + { intros T f g, + refine continuous_map.mk _ _, + { exact λ x, if h : x ∈ set.range c.inl + then f ((equiv.of_injective _ h₁.inj).symm ⟨x, h⟩) + else g ((equiv.of_injective _ h₂.inj).symm ⟨x, (this x).resolve_left h⟩) }, + rw continuous_iff_continuous_at, + intro x, + by_cases x ∈ set.range c.inl, + { revert h x, + apply (is_open.continuous_on_iff _).mp, + { rw continuous_on_iff_continuous_restrict, + convert_to continuous (f ∘ (homeomorph.of_embedding _ h₁.to_embedding).symm), + { ext ⟨x, hx⟩, exact dif_pos hx }, + continuity }, + { exact h₁.open_range } }, + { revert h x, + apply (is_open.continuous_on_iff _).mp, + { rw continuous_on_iff_continuous_restrict, + have : ∀ a, a ∉ set.range c.inl → a ∈ set.range c.inr, + { rintros a (h : a ∈ (set.range c.inl)ᶜ), rwa eq_compl_iff_is_compl.mpr h₃.symm }, + convert_to continuous + (g ∘ (homeomorph.of_embedding _ h₂.to_embedding).symm ∘ subtype.map _ this), + { ext ⟨x, hx⟩, exact dif_neg hx }, + continuity, + rw embedding_subtype_coe.to_inducing.continuous_iff, + exact continuous_subtype_coe }, + { change is_open (set.range c.inl)ᶜ, rw ← eq_compl_iff_is_compl.mpr h₃.symm, + exact h₂.open_range } } }, + { intros T f g, ext x, refine (dif_pos _).trans _, { exact ⟨x, rfl⟩ }, + { rw equiv.of_injective_symm_apply } }, + { intros T f g, ext x, refine (dif_neg _).trans _, + { rintro ⟨y, e⟩, have : c.inr x ∈ set.range c.inl ⊓ set.range c.inr := ⟨⟨_, e⟩, ⟨_, rfl⟩⟩, + rwa disjoint_iff.mp h₃.1 at this }, + { exact congr_arg g (equiv.of_injective_symm_apply _ _) } }, + { rintro T _ _ m rfl rfl, ext x, change m x = dite _ _ _, + split_ifs; exact congr_arg _ (equiv.apply_of_injective_symm _ ⟨_, _⟩).symm } } +end + +--TODO: Add analogous constructions for `pushout`. + +end Top diff --git a/src/topology/category/Top/limits/pullbacks.lean b/src/topology/category/Top/limits/pullbacks.lean new file mode 100644 index 0000000000000..56362bd553478 --- /dev/null +++ b/src/topology/category/Top/limits/pullbacks.lean @@ -0,0 +1,400 @@ +/- +Copyright (c) 2017 Scott Morrison. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Patrick Massot, Scott Morrison, Mario Carneiro, Andrew Yang +-/ +import topology.category.Top.limits.products +import category_theory.concrete_category.elementwise + +/-! +# Pullbacks in the category of topological spaces. + +-/ + +open topological_space +open category_theory +open category_theory.limits + +universes u v w + +noncomputable theory + +namespace Top + +variables {J : Type v} [small_category J] + +section pullback + +variables {X Y Z : Top.{u}} + +/-- The first projection from the pullback. -/ +abbreviation pullback_fst (f : X ⟶ Z) (g : Y ⟶ Z) : Top.of { p : X × Y // f p.1 = g p.2 } ⟶ X := +⟨prod.fst ∘ subtype.val⟩ + +/-- The second projection from the pullback. -/ +abbreviation pullback_snd (f : X ⟶ Z) (g : Y ⟶ Z) : Top.of { p : X × Y // f p.1 = g p.2 } ⟶ Y := +⟨prod.snd ∘ subtype.val⟩ + +/-- The explicit pullback cone of `X, Y` given by `{ p : X × Y // f p.1 = g p.2 }`. -/ +def pullback_cone (f : X ⟶ Z) (g : Y ⟶ Z) : pullback_cone f g := +pullback_cone.mk (pullback_fst f g) (pullback_snd f g) (by { ext ⟨x, h⟩, simp [h] }) + +/-- The constructed cone is a limit. -/ +def pullback_cone_is_limit (f : X ⟶ Z) (g : Y ⟶ Z) : + is_limit (pullback_cone f g) := pullback_cone.is_limit_aux' _ +begin + intro s, + split, swap, + exact { to_fun := λ x, ⟨⟨s.fst x, s.snd x⟩, + by simpa using concrete_category.congr_hom s.condition x⟩ }, + refine ⟨_,_,_⟩, + { ext, delta pullback_cone, simp }, + { ext, delta pullback_cone, simp }, + { intros m h₁ h₂, + ext x, + { simpa using concrete_category.congr_hom h₁ x }, + { simpa using concrete_category.congr_hom h₂ x } } +end + +/-- The pullback of two maps can be identified as a subspace of `X × Y`. -/ +def pullback_iso_prod_subtype (f : X ⟶ Z) (g : Y ⟶ Z) : + pullback f g ≅ Top.of { p : X × Y // f p.1 = g p.2 } := +(limit.is_limit _).cone_point_unique_up_to_iso (pullback_cone_is_limit f g) + +@[simp, reassoc] lemma pullback_iso_prod_subtype_inv_fst (f : X ⟶ Z) (g : Y ⟶ Z) : + (pullback_iso_prod_subtype f g).inv ≫ pullback.fst = pullback_fst f g := +by simpa [pullback_iso_prod_subtype] + +@[simp] lemma pullback_iso_prod_subtype_inv_fst_apply (f : X ⟶ Z) (g : Y ⟶ Z) + (x : { p : X × Y // f p.1 = g p.2 }) : + (pullback.fst : pullback f g ⟶ _) ((pullback_iso_prod_subtype f g).inv x) = (x : X × Y).fst := +concrete_category.congr_hom (pullback_iso_prod_subtype_inv_fst f g) x + +@[simp, reassoc] lemma pullback_iso_prod_subtype_inv_snd (f : X ⟶ Z) (g : Y ⟶ Z) : + (pullback_iso_prod_subtype f g).inv ≫ pullback.snd = pullback_snd f g := +by simpa [pullback_iso_prod_subtype] + +@[simp] lemma pullback_iso_prod_subtype_inv_snd_apply (f : X ⟶ Z) (g : Y ⟶ Z) + (x : { p : X × Y // f p.1 = g p.2 }) : + (pullback.snd : pullback f g ⟶ _) ((pullback_iso_prod_subtype f g).inv x) = (x : X × Y).snd := +concrete_category.congr_hom (pullback_iso_prod_subtype_inv_snd f g) x + +lemma pullback_iso_prod_subtype_hom_fst (f : X ⟶ Z) (g : Y ⟶ Z) : + (pullback_iso_prod_subtype f g).hom ≫ pullback_fst f g = pullback.fst := +by rw [←iso.eq_inv_comp, pullback_iso_prod_subtype_inv_fst] + +lemma pullback_iso_prod_subtype_hom_snd (f : X ⟶ Z) (g : Y ⟶ Z) : + (pullback_iso_prod_subtype f g).hom ≫ pullback_snd f g = pullback.snd := +by rw [←iso.eq_inv_comp, pullback_iso_prod_subtype_inv_snd] + +@[simp] lemma pullback_iso_prod_subtype_hom_apply {f : X ⟶ Z} {g : Y ⟶ Z} + (x : pullback f g) : (pullback_iso_prod_subtype f g).hom x = + ⟨⟨(pullback.fst : pullback f g ⟶ _) x, (pullback.snd : pullback f g ⟶ _) x⟩, + by simpa using concrete_category.congr_hom pullback.condition x⟩ := +begin + ext, + exacts [concrete_category.congr_hom (pullback_iso_prod_subtype_hom_fst f g) x, + concrete_category.congr_hom (pullback_iso_prod_subtype_hom_snd f g) x] +end + +lemma pullback_topology {X Y Z : Top.{u}} (f : X ⟶ Z) (g : Y ⟶ Z) : + (pullback f g).topological_space = + induced (pullback.fst : pullback f g ⟶ _) X.topological_space ⊓ + induced (pullback.snd : pullback f g ⟶ _) Y.topological_space := +begin + let homeo := homeo_of_iso (pullback_iso_prod_subtype f g), + refine homeo.inducing.induced.trans _, + change induced homeo (induced _ (_ ⊓ _)) = _, + simpa [induced_compose] +end + +lemma range_pullback_to_prod {X Y Z : Top} (f : X ⟶ Z) (g : Y ⟶ Z) : + set.range (prod.lift pullback.fst pullback.snd : pullback f g ⟶ X ⨯ Y) = + { x | (limits.prod.fst ≫ f) x = (limits.prod.snd ≫ g) x } := +begin + ext x, + split, + { rintros ⟨y, rfl⟩, + simp only [←comp_apply, set.mem_set_of_eq], + congr' 1, + simp [pullback.condition] }, + { intro h, + use (pullback_iso_prod_subtype f g).inv ⟨⟨_, _⟩, h⟩, + apply concrete.limit_ext, + rintro ⟨⟨⟩⟩; simp, } +end + +lemma inducing_pullback_to_prod {X Y Z : Top} (f : X ⟶ Z) (g : Y ⟶ Z) : + inducing ⇑(prod.lift pullback.fst pullback.snd : pullback f g ⟶ X ⨯ Y) := +⟨by simp [prod_topology, pullback_topology, induced_compose, ←coe_comp]⟩ + +lemma embedding_pullback_to_prod {X Y Z : Top} (f : X ⟶ Z) (g : Y ⟶ Z) : + embedding ⇑(prod.lift pullback.fst pullback.snd : pullback f g ⟶ X ⨯ Y) := +⟨inducing_pullback_to_prod f g, (Top.mono_iff_injective _).mp infer_instance⟩ + +/-- If the map `S ⟶ T` is mono, then there is a description of the image of `W ×ₛ X ⟶ Y ×ₜ Z`. -/ +lemma range_pullback_map {W X Y Z S T : Top} (f₁ : W ⟶ S) (f₂ : X ⟶ S) + (g₁ : Y ⟶ T) (g₂ : Z ⟶ T) (i₁ : W ⟶ Y) (i₂ : X ⟶ Z) (i₃ : S ⟶ T) [H₃ : mono i₃] + (eq₁ : f₁ ≫ i₃ = i₁ ≫ g₁) (eq₂ : f₂ ≫ i₃ = i₂ ≫ g₂) : + set.range (pullback.map f₁ f₂ g₁ g₂ i₁ i₂ i₃ eq₁ eq₂) = + (pullback.fst : pullback g₁ g₂ ⟶ _) ⁻¹' (set.range i₁) ∩ + (pullback.snd : pullback g₁ g₂ ⟶ _) ⁻¹' (set.range i₂) := +begin + ext, + split, + { rintro ⟨y, rfl⟩, simp, }, + rintros ⟨⟨x₁, hx₁⟩, ⟨x₂, hx₂⟩⟩, + have : f₁ x₁ = f₂ x₂, + { apply (Top.mono_iff_injective _).mp H₃, + simp only [←comp_apply, eq₁, eq₂], + simp only [comp_apply, hx₁, hx₂], + simp only [←comp_apply, pullback.condition] }, + use (pullback_iso_prod_subtype f₁ f₂).inv ⟨⟨x₁, x₂⟩, this⟩, + apply concrete.limit_ext, + rintros (_|_|_), + { simp only [Top.comp_app, limit.lift_π_apply, category.assoc, pullback_cone.mk_π_app_one, + hx₁, pullback_iso_prod_subtype_inv_fst_apply, subtype.coe_mk], + simp only [← comp_apply], + congr, + apply limit.w _ walking_cospan.hom.inl }, + { simp [hx₁] }, + { simp [hx₂] }, +end + +lemma pullback_fst_range {X Y S : Top} (f : X ⟶ S) (g : Y ⟶ S) : + set.range (pullback.fst : pullback f g ⟶ _) = { x : X | ∃ y : Y, f x = g y} := +begin + ext x, + split, + { rintro ⟨y, rfl⟩, + use (pullback.snd : pullback f g ⟶ _) y, + exact concrete_category.congr_hom pullback.condition y }, + { rintro ⟨y, eq⟩, + use (Top.pullback_iso_prod_subtype f g).inv ⟨⟨x, y⟩, eq⟩, + simp }, +end + +lemma pullback_snd_range {X Y S : Top} (f : X ⟶ S) (g : Y ⟶ S) : + set.range (pullback.snd : pullback f g ⟶ _) = { y : Y | ∃ x : X, f x = g y} := +begin + ext y, + split, + { rintro ⟨x, rfl⟩, + use (pullback.fst : pullback f g ⟶ _) x, + exact concrete_category.congr_hom pullback.condition x }, + { rintro ⟨x, eq⟩, + use (Top.pullback_iso_prod_subtype f g).inv ⟨⟨x, y⟩, eq⟩, + simp }, +end + +/-- +If there is a diagram where the morphisms `W ⟶ Y` and `X ⟶ Z` are embeddings, +then the induced morphism `W ×ₛ X ⟶ Y ×ₜ Z` is also an embedding. + + W ⟶ Y + ↘ ↘ + S ⟶ T + ↗ ↗ + X ⟶ Z +-/ +lemma pullback_map_embedding_of_embeddings {W X Y Z S T : Top} + (f₁ : W ⟶ S) (f₂ : X ⟶ S) (g₁ : Y ⟶ T) (g₂ : Z ⟶ T) {i₁ : W ⟶ Y} {i₂ : X ⟶ Z} + (H₁ : embedding i₁) (H₂ : embedding i₂) (i₃ : S ⟶ T) + (eq₁ : f₁ ≫ i₃ = i₁ ≫ g₁) (eq₂ : f₂ ≫ i₃ = i₂ ≫ g₂) : + embedding (pullback.map f₁ f₂ g₁ g₂ i₁ i₂ i₃ eq₁ eq₂) := +begin + refine embedding_of_embedding_compose (continuous_map.continuous_to_fun _) + (show continuous (prod.lift pullback.fst pullback.snd : pullback g₁ g₂ ⟶ Y ⨯ Z), from + continuous_map.continuous_to_fun _) _, + suffices : embedding + (prod.lift pullback.fst pullback.snd ≫ limits.prod.map i₁ i₂ : pullback f₁ f₂ ⟶ _), + { simpa [←coe_comp] using this }, + rw coe_comp, + refine embedding.comp (embedding_prod_map H₁ H₂) + (embedding_pullback_to_prod _ _) +end + +/-- +If there is a diagram where the morphisms `W ⟶ Y` and `X ⟶ Z` are open embeddings, and `S ⟶ T` +is mono, then the induced morphism `W ×ₛ X ⟶ Y ×ₜ Z` is also an open embedding. + W ⟶ Y + ↘ ↘ + S ⟶ T + ↗ ↗ + X ⟶ Z +-/ +lemma pullback_map_open_embedding_of_open_embeddings {W X Y Z S T : Top} + (f₁ : W ⟶ S) (f₂ : X ⟶ S) (g₁ : Y ⟶ T) (g₂ : Z ⟶ T) {i₁ : W ⟶ Y} {i₂ : X ⟶ Z} + (H₁ : open_embedding i₁) (H₂ : open_embedding i₂) (i₃ : S ⟶ T) [H₃ : mono i₃] + (eq₁ : f₁ ≫ i₃ = i₁ ≫ g₁) (eq₂ : f₂ ≫ i₃ = i₂ ≫ g₂) : + open_embedding (pullback.map f₁ f₂ g₁ g₂ i₁ i₂ i₃ eq₁ eq₂) := +begin + split, + { apply pullback_map_embedding_of_embeddings + f₁ f₂ g₁ g₂ H₁.to_embedding H₂.to_embedding i₃ eq₁ eq₂ }, + { rw range_pullback_map, + apply is_open.inter; apply continuous.is_open_preimage, + continuity, + exacts [H₁.open_range, H₂.open_range] } +end + +lemma snd_embedding_of_left_embedding {X Y S : Top} + {f : X ⟶ S} (H : embedding f) (g : Y ⟶ S) : + embedding ⇑(pullback.snd : pullback f g ⟶ Y) := +begin + convert (homeo_of_iso (as_iso (pullback.snd : pullback (𝟙 S) g ⟶ _))).embedding.comp + (pullback_map_embedding_of_embeddings f g (𝟙 _) g H + (homeo_of_iso (iso.refl _)).embedding (𝟙 _) rfl (by simp)), + erw ←coe_comp, + simp +end + +lemma fst_embedding_of_right_embedding {X Y S : Top} + (f : X ⟶ S) {g : Y ⟶ S} (H : embedding g) : + embedding ⇑(pullback.fst : pullback f g ⟶ X) := +begin + convert (homeo_of_iso (as_iso (pullback.fst : pullback f (𝟙 S) ⟶ _))).embedding.comp + (pullback_map_embedding_of_embeddings f g f (𝟙 _) + (homeo_of_iso (iso.refl _)).embedding H (𝟙 _) rfl (by simp)), + erw ←coe_comp, + simp +end + +lemma embedding_of_pullback_embeddings {X Y S : Top} + {f : X ⟶ S} {g : Y ⟶ S} (H₁ : embedding f) (H₂ : embedding g) : + embedding (limit.π (cospan f g) walking_cospan.one) := +begin + convert H₂.comp (snd_embedding_of_left_embedding H₁ g), + erw ←coe_comp, + congr, + exact (limit.w _ walking_cospan.hom.inr).symm +end + +lemma snd_open_embedding_of_left_open_embedding {X Y S : Top} + {f : X ⟶ S} (H : open_embedding f) (g : Y ⟶ S) : + open_embedding ⇑(pullback.snd : pullback f g ⟶ Y) := +begin + convert (homeo_of_iso (as_iso (pullback.snd : pullback (𝟙 S) g ⟶ _))).open_embedding.comp + (pullback_map_open_embedding_of_open_embeddings f g (𝟙 _) g H + (homeo_of_iso (iso.refl _)).open_embedding (𝟙 _) rfl (by simp)), + erw ←coe_comp, + simp +end + +lemma fst_open_embedding_of_right_open_embedding {X Y S : Top} + (f : X ⟶ S) {g : Y ⟶ S} (H : open_embedding g) : + open_embedding ⇑(pullback.fst : pullback f g ⟶ X) := +begin + convert (homeo_of_iso (as_iso (pullback.fst : pullback f (𝟙 S) ⟶ _))).open_embedding.comp + (pullback_map_open_embedding_of_open_embeddings f g f (𝟙 _) + (homeo_of_iso (iso.refl _)).open_embedding H (𝟙 _) rfl (by simp)), + erw ←coe_comp, + simp +end + +/-- If `X ⟶ S`, `Y ⟶ S` are open embeddings, then so is `X ×ₛ Y ⟶ S`. -/ +lemma open_embedding_of_pullback_open_embeddings {X Y S : Top} + {f : X ⟶ S} {g : Y ⟶ S} (H₁ : open_embedding f) (H₂ : open_embedding g) : + open_embedding (limit.π (cospan f g) walking_cospan.one) := +begin + convert H₂.comp (snd_open_embedding_of_left_open_embedding H₁ g), + erw ←coe_comp, + congr, + exact (limit.w _ walking_cospan.hom.inr).symm +end + +lemma fst_iso_of_right_embedding_range_subset {X Y S : Top} (f : X ⟶ S) {g : Y ⟶ S} + (hg : embedding g) (H : set.range f ⊆ set.range g) : is_iso (pullback.fst : pullback f g ⟶ X) := +begin + let : (pullback f g : Top) ≃ₜ X := + (homeomorph.of_embedding _ (fst_embedding_of_right_embedding f hg)).trans + { to_fun := coe, + inv_fun := (λ x, ⟨x, + by { rw pullback_fst_range, exact ⟨_, (H (set.mem_range_self x)).some_spec.symm⟩ }⟩), + left_inv := λ ⟨_,_⟩, rfl, + right_inv := λ x, rfl }, + convert is_iso.of_iso (iso_of_homeo this), + ext, + refl +end + +lemma snd_iso_of_left_embedding_range_subset {X Y S : Top} {f : X ⟶ S} (hf : embedding f) + (g : Y ⟶ S) (H : set.range g ⊆ set.range f) : is_iso (pullback.snd : pullback f g ⟶ Y) := +begin + let : (pullback f g : Top) ≃ₜ Y := + (homeomorph.of_embedding _ (snd_embedding_of_left_embedding hf g)).trans + { to_fun := coe, + inv_fun := (λ x, ⟨x, + by { rw pullback_snd_range, exact ⟨_, (H (set.mem_range_self x)).some_spec⟩ }⟩), + left_inv := λ ⟨_,_⟩, rfl, + right_inv := λ x, rfl }, + convert is_iso.of_iso (iso_of_homeo this), + ext, + refl +end + +lemma pullback_snd_image_fst_preimage (f : X ⟶ Z) (g : Y ⟶ Z) (U : set X) : + (pullback.snd : pullback f g ⟶ _) '' ((pullback.fst : pullback f g ⟶ _) ⁻¹' U) = + g ⁻¹' (f '' U) := +begin + ext x, + split, + { rintros ⟨y, hy, rfl⟩, + exact ⟨(pullback.fst : pullback f g ⟶ _) y, hy, + concrete_category.congr_hom pullback.condition y⟩ }, + { rintros ⟨y, hy, eq⟩, + exact ⟨(Top.pullback_iso_prod_subtype f g).inv ⟨⟨_,_⟩, eq⟩, by simpa, by simp⟩ }, +end + +lemma pullback_fst_image_snd_preimage (f : X ⟶ Z) (g : Y ⟶ Z) (U : set Y) : + (pullback.fst : pullback f g ⟶ _) '' ((pullback.snd : pullback f g ⟶ _) ⁻¹' U) = + f ⁻¹' (g '' U) := +begin + ext x, + split, + { rintros ⟨y, hy, rfl⟩, + exact ⟨(pullback.snd : pullback f g ⟶ _) y, hy, + (concrete_category.congr_hom pullback.condition y).symm⟩ }, + { rintros ⟨y, hy, eq⟩, + exact ⟨(Top.pullback_iso_prod_subtype f g).inv ⟨⟨_,_⟩,eq.symm⟩, by simpa, by simp⟩ }, +end + +end pullback + + + +lemma coinduced_of_is_colimit {F : J ⥤ Top.{max v u}} (c : cocone F) (hc : is_colimit c) : + c.X.topological_space = ⨆ j, (F.obj j).topological_space.coinduced (c.ι.app j) := +begin + let homeo := homeo_of_iso (hc.cocone_point_unique_up_to_iso (colimit_cocone_is_colimit F)), + ext, + refine homeo.symm.is_open_preimage.symm.trans (iff.trans _ is_open_supr_iff.symm), + exact is_open_supr_iff +end + +lemma colimit_topology (F : J ⥤ Top.{max v u}) : + (colimit F).topological_space = ⨆ j, (F.obj j).topological_space.coinduced (colimit.ι F j) := +coinduced_of_is_colimit _ (colimit.is_colimit F) + +lemma colimit_is_open_iff (F : J ⥤ Top.{max v u}) (U : set ((colimit F : _) : Type (max v u))) : + is_open U ↔ ∀ j, is_open (colimit.ι F j ⁻¹' U) := +begin + conv_lhs { rw colimit_topology F }, + exact is_open_supr_iff +end + +lemma coequalizer_is_open_iff (F : walking_parallel_pair ⥤ Top.{u}) + (U : set ((colimit F : _) : Type u)) : + is_open U ↔ is_open (colimit.ι F walking_parallel_pair.one ⁻¹' U) := +begin + rw colimit_is_open_iff.{u}, + split, + { intro H, exact H _ }, + { intros H j, + cases j, + { rw ←colimit.w F walking_parallel_pair_hom.left, + exact (F.map walking_parallel_pair_hom.left).continuous_to_fun.is_open_preimage _ H }, + { exact H } } +end + +end Top diff --git a/src/topology/gluing.lean b/src/topology/gluing.lean index b254ae5a0ba85..6186f6da75bf7 100644 --- a/src/topology/gluing.lean +++ b/src/topology/gluing.lean @@ -5,7 +5,7 @@ Authors: Andrew Yang -/ import category_theory.glue_data import category_theory.concrete_category.elementwise -import topology.category.Top.limits +import topology.category.Top.limits.pullbacks import topology.category.Top.opens /-! From cdb01be3c499930fd29be05dce960f4d8d201c54 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 26 Apr 2023 18:21:07 +0000 Subject: [PATCH 0881/1291] =?UTF-8?q?feat(algebra/big=5Foperators/fin):=20?= =?UTF-8?q?The=20equivalence=20between=20`fin=20n=20=E2=86=92=20fin=20m`?= =?UTF-8?q?=20and=20`fin=20(m=20^=20n)`=20(#14817)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Eric Wieser --- src/algebra/big_operators/fin.lean | 122 +++++++++++++++++++++++++++++ 1 file changed, 122 insertions(+) diff --git a/src/algebra/big_operators/fin.lean b/src/algebra/big_operators/fin.lean index a6d64ed604a40..d0182e6396d4e 100644 --- a/src/algebra/big_operators/fin.lean +++ b/src/algebra/big_operators/fin.lean @@ -20,6 +20,9 @@ The most important results are the induction formulas `fin.prod_univ_cast_succ` and `fin.prod_univ_succ`, and the formula `fin.prod_const` for the product of a constant function. These results have variants for sums instead of products. +## Main declarations + +* `fin_function_fin_equiv`: An explicit equivalence between `fin n → fin m` and `fin (m ^ n)`. -/ open_locale big_operators @@ -209,6 +212,125 @@ end partial_prod end fin +/-- Equivalence between `fin n → fin m` and `fin (m ^ n)`. -/ +@[simps] def fin_function_fin_equiv {m n : ℕ} : (fin n → fin m) ≃ fin (m ^ n) := +equiv.of_right_inverse_of_card_le + (le_of_eq $ by simp_rw [fintype.card_fun, fintype.card_fin]) + (λ f, ⟨∑ i, f i * m ^ (i : ℕ), begin + induction n with n ih generalizing f, + { simp }, + cases m, + { exact is_empty_elim (f $ fin.last _) }, + simp_rw [fin.sum_univ_cast_succ, fin.coe_cast_succ, fin.coe_last], + refine (add_lt_add_of_lt_of_le (ih _) $ mul_le_mul_right' (fin.is_le _) _).trans_eq _, + rw [←one_add_mul, add_comm, pow_succ], + end⟩) + (λ a b, ⟨a / m ^ (b : ℕ) % m, begin + cases n, + { exact b.elim0 }, + cases m, + { rw zero_pow n.succ_pos at a, + exact a.elim0 }, + { exact nat.mod_lt _ m.succ_pos } + end⟩) + (λ a, begin + dsimp, + induction n with n ih generalizing a, + { haveI : subsingleton (fin (m ^ 0)) := (fin.cast $ pow_zero _).to_equiv.subsingleton, + exact subsingleton.elim _ _ }, + simp_rw [fin.forall_iff, fin.ext_iff, fin.coe_mk] at ih, + ext, + simp_rw [fin.coe_mk, fin.sum_univ_succ, fin.coe_zero, fin.coe_succ, pow_zero, nat.div_one, + mul_one, pow_succ, ←nat.div_div_eq_div_mul, mul_left_comm _ m, ←mul_sum], + rw [ih _ (nat.div_lt_of_lt_mul a.is_lt), nat.mod_add_div], + end) + +lemma fin_function_fin_equiv_apply {m n : ℕ} (f : fin n → fin m): + (fin_function_fin_equiv f : ℕ) = ∑ (i : fin n), ↑(f i) * m ^ (i : ℕ) := rfl + +lemma fin_function_fin_equiv_single {m n : ℕ} [ne_zero m] (i : fin n) (j : fin m) : + (fin_function_fin_equiv (pi.single i j) : ℕ) = j * m ^ (i : ℕ) := +begin + rw [fin_function_fin_equiv_apply, fintype.sum_eq_single i, pi.single_eq_same], + rintro x hx, + rw [pi.single_eq_of_ne hx, fin.coe_zero, zero_mul], +end + +/-- Equivalence between `Π i : fin m, fin (n i)` and `fin (∏ i : fin m, n i)`. -/ +def fin_pi_fin_equiv {m : ℕ} {n : fin m → ℕ} : + (Π i : fin m, fin (n i)) ≃ fin (∏ i : fin m, n i) := +equiv.of_right_inverse_of_card_le + (le_of_eq $ by simp_rw [fintype.card_pi, fintype.card_fin]) + (λ f, ⟨∑ i, f i * ∏ j, n (fin.cast_le i.is_lt.le j), begin + induction m with m ih generalizing f, + { simp }, + rw [fin.prod_univ_cast_succ, fin.sum_univ_cast_succ], + suffices : ∀ (n : fin m → ℕ) (nn : ℕ) (f : Π i : fin m, fin (n i)) (fn : fin nn), + ∑ i : fin m, ↑(f i) * ∏ j : fin i, n (fin.cast_le i.prop.le j) + ↑fn * ∏ j, n j + < (∏ i : fin m, n i) * nn, + { replace this := this (fin.init n) (n (fin.last _)) (fin.init f) (f (fin.last _)), + rw ←fin.snoc_init_self f, + simp only [←fin.snoc_init_self n] { single_pass := tt }, + simp_rw [fin.snoc_cast_succ, fin.init_snoc, fin.snoc_last, fin.snoc_init_self n], + exact this }, + intros n nn f fn, + cases nn, + { exact is_empty_elim fn }, + refine (add_lt_add_of_lt_of_le (ih _) $ mul_le_mul_right' (fin.is_le _) _).trans_eq _, + rw [←one_add_mul, mul_comm, add_comm], + end⟩) + (λ a b, ⟨a / (∏ j : fin b, n (fin.cast_le b.is_lt.le j)) % n b, begin + cases m, + { exact b.elim0 }, + cases h : n b with nb, + { rw prod_eq_zero (finset.mem_univ _) h at a, + exact is_empty_elim a }, + exact nat.mod_lt _ nb.succ_pos, + end⟩) + (begin + intro a, revert a, dsimp only [fin.coe_mk], + refine fin.cons_induction _ _ n, + { intro a, + haveI : subsingleton (fin (∏ i : fin 0, i.elim0)) := + (fin.cast $ prod_empty).to_equiv.subsingleton, + exact subsingleton.elim _ _ }, + { intros n x xs ih a, + simp_rw [fin.forall_iff, fin.ext_iff, fin.coe_mk] at ih, + ext, + simp_rw [fin.coe_mk, fin.sum_univ_succ, fin.cons_succ], + have := λ i : fin n, fintype.prod_equiv (fin.cast $ fin.coe_succ i).to_equiv + (λ j, (fin.cons x xs : _ → ℕ) (fin.cast_le (fin.is_lt _).le j)) + (λ j, (fin.cons x xs : _ → ℕ) (fin.cast_le (nat.succ_le_succ (fin.is_lt _).le) j)) + (λ j, rfl), + simp_rw [this], clear this, + dsimp only [fin.coe_zero], + simp_rw [fintype.prod_empty, nat.div_one, mul_one, fin.cons_zero, fin.prod_univ_succ], + change _ + ∑ y : _, ((_ / (x * _)) % _) * (x * _) = _, + simp_rw [←nat.div_div_eq_div_mul, mul_left_comm (_ % _ : ℕ), ←mul_sum], + convert nat.mod_add_div _ _, + refine eq.trans _ (ih (a / x) (nat.div_lt_of_lt_mul $ a.is_lt.trans_eq _)), + swap, + { convert fin.prod_univ_succ (fin.cons x xs : (Π _, ℕ)), + simp_rw fin.cons_succ }, + congr' with i, + congr' with j, + { cases j, refl }, + { cases j, refl } } + end) + +lemma fin_pi_fin_equiv_apply {m : ℕ} {n : fin m → ℕ} (f : Π i : fin m, fin (n i)): + (fin_pi_fin_equiv f : ℕ) = ∑ i, f i * ∏ j, n (fin.cast_le i.is_lt.le j) := rfl + +lemma fin_pi_fin_equiv_single {m : ℕ} {n : fin m → ℕ} [Π i, ne_zero (n i)] + (i : fin m) (j : fin (n i)) : + (fin_pi_fin_equiv (pi.single i j : Π i : fin m, fin (n i)) : ℕ) + = j * ∏ j, n (fin.cast_le i.is_lt.le j) := +begin + rw [fin_pi_fin_equiv_apply, fintype.sum_eq_single i, pi.single_eq_same], + rintro x hx, + rw [pi.single_eq_of_ne hx, fin.coe_zero, zero_mul], +end + namespace list section comm_monoid From dc7ac07acd84584426773e69e51035bea9a770e7 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 26 Apr 2023 18:21:08 +0000 Subject: [PATCH 0882/1291] chore(algebra/star/basic): generalize quaternion lemmas (#18802) We already had most of these lemmas specialized to quaternions; this generalizes them to any star ring. We should consider replacing `quaternion.conj` with `star` in future. --- src/algebra/quaternion.lean | 29 +++++++++++++---------------- src/algebra/star/basic.lean | 26 ++++++++++++++++++++++++++ 2 files changed, 39 insertions(+), 16 deletions(-) diff --git a/src/algebra/quaternion.lean b/src/algebra/quaternion.lean index df53f8c4b7f83..8cc429e41dca3 100644 --- a/src/algebra/quaternion.lean +++ b/src/algebra/quaternion.lean @@ -362,11 +362,17 @@ lemma conj_add : (a + b).conj = a.conj + b.conj := conj.map_add a b @[simp] lemma conj_mul : (a * b).conj = b.conj * a.conj := by ext; simp; ring_exp -lemma conj_conj_mul : (a.conj * b).conj = b.conj * a := -by rw [conj_mul, conj_conj] +instance : star_ring ℍ[R, c₁, c₂] := +{ star := conj, + star_involutive := conj_conj, + star_add := conj_add, + star_mul := conj_mul } + +@[simp] lemma star_def (a : ℍ[R, c₁, c₂]) : star a = conj a := rfl + +lemma conj_conj_mul : (a.conj * b).conj = b.conj * a := star_star_mul _ _ -lemma conj_mul_conj : (a * b.conj).conj = b * a.conj := -by rw [conj_mul, conj_conj] +lemma conj_mul_conj : (a * b.conj).conj = b * a.conj := star_mul_star _ _ lemma self_add_conj' : a + a.conj = ↑(2 * a.re) := by ext; simp [two_mul] @@ -389,18 +395,16 @@ lemma commute_self_conj : commute a a.conj := a.commute_conj_self.symm lemma commute_conj_conj {a b : ℍ[R, c₁, c₂]} (h : commute a b) : commute a.conj b.conj := -calc a.conj * b.conj = (b * a).conj : (conj_mul b a).symm - ... = (a * b).conj : by rw h.eq - ... = b.conj * a.conj : conj_mul a b +h.star_star @[simp, norm_cast] lemma conj_coe : conj (x : ℍ[R, c₁, c₂]) = x := by ext; simp @[simp] lemma conj_im : conj a.im = - a.im := im_conj _ @[simp, norm_cast] lemma conj_nat_cast (n : ℕ) : conj (n : ℍ[R, c₁, c₂]) = n := -by rw [←coe_nat_cast, conj_coe] +@star_nat_cast ℍ[R, c₁, c₂] _ _ n @[simp, norm_cast] lemma conj_int_cast (z : ℤ) : conj (z : ℍ[R, c₁, c₂]) = z := -by rw [←coe_int_cast, conj_coe] +@star_int_cast ℍ[R, c₁, c₂] _ _ z @[simp] lemma conj_smul [monoid S] [distrib_mul_action S R] (s : S) (a : ℍ[R, c₁, c₂]) : conj (s • a) = s • conj a := @@ -441,13 +445,6 @@ lemma conj_neg : (-a).conj = -a.conj := (conj : ℍ[R, c₁, c₂] ≃ₗ[R] _). lemma conj_sub : (a - b).conj = a.conj - b.conj := (conj : ℍ[R, c₁, c₂] ≃ₗ[R] _).map_sub a b -instance : star_ring ℍ[R, c₁, c₂] := -{ star := conj, - star_involutive := conj_conj, - star_add := conj_add, - star_mul := conj_mul } - -@[simp] lemma star_def (a : ℍ[R, c₁, c₂]) : star a = conj a := rfl @[simp] lemma conj_pow (n : ℕ) : (a ^ n).conj = a.conj ^ n := star_pow _ _ diff --git a/src/algebra/star/basic.lean b/src/algebra/star/basic.lean index de34a4d606ac6..9b18517de2d79 100644 --- a/src/algebra/star/basic.lean +++ b/src/algebra/star/basic.lean @@ -94,6 +94,9 @@ star_involutive _ lemma star_injective [has_involutive_star R] : function.injective (star : R → R) := star_involutive.injective +@[simp] lemma star_inj [has_involutive_star R] {x y : R} : star x = star y ↔ x = y := +star_injective.eq_iff + /-- `star` as an equivalence when it is involutive. -/ protected def equiv.star [has_involutive_star R] : equiv.perm R := star_involutive.to_perm _ @@ -126,6 +129,29 @@ class star_semigroup (R : Type u) [semigroup R] extends has_involutive_star R := export star_semigroup (star_mul) attribute [simp] star_mul +section star_semigroup +variables [semigroup R] [star_semigroup R] + +lemma star_star_mul (x y : R) : star (star x * y) = star y * x := by rw [star_mul, star_star] + +lemma star_mul_star (x y : R) : star (x * star y) = y * star x := by rw [star_mul, star_star] + +@[simp] lemma semiconj_by_star_star_star {x y z : R} : + semiconj_by (star x) (star z) (star y) ↔ semiconj_by x y z := +by simp_rw [semiconj_by, ←star_mul, star_inj, eq_comm] + +alias semiconj_by_star_star_star ↔ _ semiconj_by.star_star_star + +@[simp] lemma commute_star_star {x y : R} : commute (star x) (star y) ↔ commute x y := +semiconj_by_star_star_star + +alias commute_star_star ↔ _ commute.star_star + +lemma commute_star_comm {x y : R} : commute (star x) y ↔ commute x (star y) := +by rw [←commute_star_star, star_star] + +end star_semigroup + /-- In a commutative ring, make `simp` prefer leaving the order unchanged. -/ @[simp] lemma star_mul' [comm_semigroup R] [star_semigroup R] (x y : R) : star (x * y) = star x * star y := From 9b2b58d6b14b895b2f375108e765cb47de71aebd Mon Sep 17 00:00:00 2001 From: Felix-Weilacher Date: Wed, 26 Apr 2023 18:21:10 +0000 Subject: [PATCH 0883/1291] feat(measure_theory/constructions): add the Borel isomorphism theorem. (#18864) Add a file `measure_theory/constructions/borel_isomorphism.lean` with several versions of the Borel isomorphism theorem. Also add fairly small supporting lemmas in several other files. Most notable here is the addition of a type class for measurable_spaces whose sigma-algebras are generated by some countable set, and a theorem that such spaces admit measurable injections to the Cantor space. Co-authored-by: Felix-Weilacher <112423742+Felix-Weilacher@users.noreply.github.com> --- .../constructions/borel_space.lean | 10 +++ src/measure_theory/constructions/polish.lean | 78 ++++++++++++++++++- src/measure_theory/measurable_space.lean | 64 +++++++++++++++ src/topology/perfect.lean | 19 ++++- 4 files changed, 164 insertions(+), 7 deletions(-) diff --git a/src/measure_theory/constructions/borel_space.lean b/src/measure_theory/constructions/borel_space.lean index 13006781fd0a4..426ffacdb82d2 100644 --- a/src/measure_theory/constructions/borel_space.lean +++ b/src/measure_theory/constructions/borel_space.lean @@ -258,6 +258,16 @@ instance subtype.opens_measurable_space {α : Type*} [topological_space α] [mea opens_measurable_space s := ⟨by { rw [borel_comap], exact comap_mono h.1 }⟩ +@[priority 100] +instance borel_space.countably_generated {α : Type*} [topological_space α] [measurable_space α] + [borel_space α] [second_countable_topology α] : countably_generated α := +begin + obtain ⟨b, bct, -, hb⟩ := exists_countable_basis α, + refine ⟨⟨b, bct, _⟩⟩, + borelize α, + exact hb.borel_eq_generate_from, +end + theorem _root_.measurable_set.induction_on_open [topological_space α] [measurable_space α] [borel_space α] {C : set α → Prop} (h_open : ∀ U, is_open U → C U) (h_compl : ∀ t, measurable_set t → C t → C tᶜ) diff --git a/src/measure_theory/constructions/polish.lean b/src/measure_theory/constructions/polish.lean index e04eadf44e41c..43bc4770d445f 100644 --- a/src/measure_theory/constructions/polish.lean +++ b/src/measure_theory/constructions/polish.lean @@ -1,9 +1,9 @@ /- Copyright (c) 2022 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Sébastien Gouëzel +Authors: Sébastien Gouëzel, Felix Weilacher -/ -import topology.metric_space.polish +import topology.perfect import measure_theory.constructions.borel_space /-! @@ -29,7 +29,7 @@ Then, we show Lusin's theorem that two disjoint analytic sets can be separated b * `analytic_set.measurably_separable` shows that two disjoint analytic sets are separated by a Borel set. -Finally, we prove the Lusin-Souslin theorem that a continuous injective image of a Borel subset of +We then prove the Lusin-Souslin theorem that a continuous injective image of a Borel subset of a Polish space is Borel. The proof of this nontrivial result relies on the above results on analytic sets. @@ -43,6 +43,11 @@ analytic sets. to a second-countable topological space is a measurable embedding. * `is_clopenable_iff_measurable_set`: in a Polish space, a set is clopenable (i.e., it can be made open and closed by using a finer Polish topology) if and only if it is Borel-measurable. + +We use this to prove several versions of the Borel isomorphism theorem. + +* `measurable_equiv_of_not_countable` : Any two uncountable Polish spaces are Borel isomorphic. +* `equiv.measurable_equiv` : Any two Polish spaces of the same cardinality are Borel. isomorphic. -/ open set function polish_space pi_nat topological_space metric filter @@ -709,3 +714,70 @@ begin end end measure_theory + +/-! ### The Borel Isomorphism Theorem -/ + +/-Note: Move to topology/metric_space/polish when porting. -/ +@[priority 50] +instance polish_of_countable [h : countable α] [discrete_topology α] : polish_space α := +begin + obtain ⟨f, hf⟩ := h.exists_injective_nat, + have : closed_embedding f, + { apply closed_embedding_of_continuous_injective_closed continuous_of_discrete_topology hf, + exact λ t _, is_closed_discrete _, }, + exact this.polish_space, +end + +/-Note: This is to avoid a loop in TC inference. When ported to Lean 4, this will not +be necessary, and `second_countable_of_polish` should probably +just be added as an instance soon after the definition of `polish_space`.-/ +private lemma second_countable_of_polish [h : polish_space α] : second_countable_topology α := +h.second_countable + +local attribute [-instance] polish_space_of_complete_second_countable +local attribute [instance] second_countable_of_polish + +namespace polish_space + +variables {β : Type*} [topological_space β] [polish_space α] [polish_space β] +variables [measurable_space α] [measurable_space β] [borel_space α] [borel_space β] + +noncomputable theory + +/-- If two Polish spaces admit Borel measurable injections to one another, +then they are Borel isomorphic.-/ +def borel_schroeder_bernstein + {f : α → β} {g : β → α} + (fmeas : measurable f) (finj : function.injective f) + (gmeas : measurable g) (ginj : function.injective g) : + α ≃ᵐ β := +(fmeas.measurable_embedding finj).schroeder_bernstein (gmeas.measurable_embedding ginj) + +/-- Any uncountable Polish space is Borel isomorphic to the Cantor space `ℕ → bool`.-/ +def measurable_equiv_nat_bool_of_not_countable (h : ¬ countable α) : α ≃ᵐ (ℕ → bool) := +begin + apply nonempty.some, + obtain ⟨f, -, fcts, finj⟩ := is_closed_univ.exists_nat_bool_injection_of_not_countable + (by rwa [← countable_coe_iff, (equiv.set.univ _).countable_iff]), + obtain ⟨g, gmeas, ginj⟩ := measurable_space.measurable_injection_cantor_of_countably_generated α, + exact ⟨borel_schroeder_bernstein gmeas ginj fcts.measurable finj⟩, +end + +/-- The **Borel Isomorphism Theorem**: Any two uncountable Polish spaces are Borel isomorphic.-/ +def measurable_equiv_of_not_countable (hα : ¬ countable α) (hβ : ¬ countable β ) : α ≃ᵐ β := +(measurable_equiv_nat_bool_of_not_countable hα).trans + (measurable_equiv_nat_bool_of_not_countable hβ).symm + +/-- The **Borel Isomorphism Theorem**: If two Polish spaces have the same cardinality, +they are Borel isomorphic.-/ +def equiv.measurable_equiv (e : α ≃ β) : α ≃ᵐ β := +begin + by_cases h : countable α, + { letI := h, + letI := countable.of_equiv α e, + use e; apply measurable_of_countable, }, + refine measurable_equiv_of_not_countable h _, + rwa e.countable_iff at h, +end + +end polish_space diff --git a/src/measure_theory/measurable_space.lean b/src/measure_theory/measurable_space.lean index 01bf4ed280315..a51e01907c94b 100644 --- a/src/measure_theory/measurable_space.lean +++ b/src/measure_theory/measurable_space.lean @@ -315,6 +315,10 @@ begin { simp only [preimage_singleton_eq_empty.2 hyf, measurable_set.empty] } end +lemma measurable_to_countable' [measurable_space α] [countable α] [measurable_space β] {f : β → α} + (h : ∀ x, measurable_set (f ⁻¹' {x})) : measurable f := +measurable_to_countable (λ y, h (f y)) + @[measurability] lemma measurable_unit [measurable_space α] (f : unit → α) : measurable f := measurable_from_top @@ -327,6 +331,15 @@ measurable_from_top lemma measurable_to_nat {f : α → ℕ} : (∀ y, measurable_set (f ⁻¹' {f y})) → measurable f := measurable_to_countable +lemma measurable_to_bool {f : α → bool} (h : measurable_set (f⁻¹' {tt})) : measurable f := +begin + apply measurable_to_countable', + rintros (-|-), + { convert h.compl, + rw [← preimage_compl, bool.compl_singleton, bool.bnot_tt] }, + exact h, +end + lemma measurable_find_greatest' {p : α → ℕ → Prop} [∀ x, decidable_pred (p x)] {N : ℕ} (hN : ∀ k ≤ N, measurable_set {x | nat.find_greatest (p x) N = k}) : measurable (λ x, nat.find_greatest (p x) N) := @@ -1364,6 +1377,57 @@ end end measurable_embedding +section countably_generated + +namespace measurable_space + +variable (α) + +/-- We say a measurable space is countably generated +if can be generated by a countable set of sets.-/ +class countably_generated [m : measurable_space α] : Prop := + (is_countably_generated : ∃ b : set (set α), b.countable ∧ m = generate_from b) + +open_locale classical + +/-- If a measurable space is countably generated, it admits a measurable injection +into the Cantor space `ℕ → bool` (equipped with the product sigma algebra). -/ +theorem measurable_injection_cantor_of_countably_generated +[measurable_space α] [h : countably_generated α] [measurable_singleton_class α] : +∃ f : α → (ℕ → bool), measurable f ∧ function.injective f := +begin + obtain ⟨b, bct, hb⟩ := h.is_countably_generated, + obtain ⟨e, he⟩ := set.countable.exists_eq_range (bct.insert ∅) (insert_nonempty _ _), + rw [← generate_from_insert_empty, he] at hb, + refine ⟨λ x n, to_bool (x ∈ e n), _, _⟩, + { rw measurable_pi_iff, + intro n, + apply measurable_to_bool, + simp only [preimage, mem_singleton_iff, to_bool_iff, set_of_mem_eq], + rw hb, + apply measurable_set_generate_from, + use n, }, + intros x y hxy, + have : ∀ s : set α, measurable_set s → (x ∈ s ↔ y ∈ s) := λ s, by + { rw hb, + apply generate_from_induction, + { rintros - ⟨n, rfl⟩, + rw ← bool.to_bool_eq, + rw funext_iff at hxy, + exact hxy n }, + { tauto }, + { intro t, + tauto }, + intros t ht, + simp_rw [mem_Union, ht], }, + specialize this {y} measurable_set_eq, + simpa only [mem_singleton, iff_true], +end + +end measurable_space + +end countably_generated + namespace filter variables [measurable_space α] diff --git a/src/topology/perfect.lean b/src/topology/perfect.lean index a92dd35cf4e0d..1321cd68d34d8 100644 --- a/src/topology/perfect.lean +++ b/src/topology/perfect.lean @@ -3,8 +3,7 @@ Copyright (c) 2022 Felix Weilacher. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Felix Weilacher -/ -import topology.separation -import topology.bases +import topology.metric_space.polish import topology.metric_space.cantor_scheme /-! @@ -223,7 +222,7 @@ end end kernel end basic -section cantor_inj +section cantor_inj_metric open function open_locale ennreal @@ -324,4 +323,16 @@ begin simpa only [← subtype.val_inj] using hdisj'.map_injective hxy, end -end cantor_inj +end cantor_inj_metric + +/-- Any closed uncountable subset of a Polish space admits a continuous injection +from the Cantor space `ℕ → bool`.-/ +theorem is_closed.exists_nat_bool_injection_of_not_countable {α : Type*} + [topological_space α] [polish_space α] {C : set α} (hC : is_closed C) (hunc : ¬ C.countable) : + ∃ f : (ℕ → bool) → α, (range f) ⊆ C ∧ continuous f ∧ function.injective f := +begin + letI := upgrade_polish_space α, + obtain ⟨D, hD, Dnonempty, hDC⟩ := exists_perfect_nonempty_of_is_closed_of_not_countable hC hunc, + obtain ⟨f, hfD, hf⟩ := hD.exists_nat_bool_injection Dnonempty, + exact ⟨f, hfD.trans hDC, hf⟩, +end From 57911c5a05a1b040598e1e15b189f035ac5cc33c Mon Sep 17 00:00:00 2001 From: Eric Rodriguez <37984851+ericrbg@users.noreply.github.com> Date: Wed, 26 Apr 2023 20:55:02 +0000 Subject: [PATCH 0884/1291] chore(finsupp/basic): weaken hypotheses (#18874) these are already in mathlib4, I just missed them in mathlib3. --- src/data/finsupp/basic.lean | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/src/data/finsupp/basic.lean b/src/data/finsupp/basic.lean index de49e4b2009a1..7d90e3d52214d 100644 --- a/src/data/finsupp/basic.lean +++ b/src/data/finsupp/basic.lean @@ -1259,16 +1259,16 @@ Throughout this section, some `monoid` and `semiring` arguments are specified wi `[]`. See note [implicit instance arguments]. -/ -@[simp] lemma coe_smul [add_monoid M] [smul_zero_class R M] +@[simp] lemma coe_smul [has_zero M] [smul_zero_class R M] (b : R) (v : α →₀ M) : ⇑(b • v) = b • v := rfl -lemma smul_apply [add_monoid M] [smul_zero_class R M] +lemma smul_apply [has_zero M] [smul_zero_class R M] (b : R) (v : α →₀ M) (a : α) : (b • v) a = b • (v a) := rfl -lemma _root_.is_smul_regular.finsupp [add_monoid M] [smul_zero_class R M] {k : R} +lemma _root_.is_smul_regular.finsupp [has_zero M] [smul_zero_class R M] {k : R} (hk : is_smul_regular M k) : is_smul_regular (α →₀ M) k := λ _ _ h, ext $ λ i, hk (congr_fun h i) -instance [nonempty α] [add_monoid M] [smul_zero_class R M] [has_faithful_smul R M] : +instance [nonempty α] [has_zero M] [smul_zero_class R M] [has_faithful_smul R M] : has_faithful_smul R (α →₀ M) := { eq_of_smul_eq_smul := λ r₁ r₂ h, let ⟨a⟩ := ‹nonempty α› in eq_of_smul_eq_smul $ λ m : M, by simpa using congr_fun (h (single a m)) a } From 9d2f0748e6c50d7a2657c564b1ff2c695b39148d Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Thu, 27 Apr 2023 05:18:06 +0000 Subject: [PATCH 0885/1291] chore(*): add mathlib4 synchronization comments (#18853) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.category.Module.tannaka` * `algebra.homology.homotopy_category` * `algebra.homology.short_exact.preadditive` * `algebraic_topology.dold_kan.equivalence_additive` * `algebraic_topology.dold_kan.functor_gamma` * `algebraic_topology.dold_kan.gamma_comp_n` * `algebraic_topology.dold_kan.homotopy_equivalence` * `algebraic_topology.dold_kan.n_comp_gamma` * `algebraic_topology.dold_kan.normalized` * `analysis.box_integral.partition.basic` * `analysis.box_integral.partition.split` * `analysis.convex.caratheodory` * `analysis.convex.combination` * `analysis.convex.independent` * `analysis.convex.jensen` * `analysis.convex.join` * `analysis.convex.normed` * `analysis.convex.simplicial_complex.basic` * `analysis.convex.topology` * `analysis.locally_convex.bounded` * `analysis.normed_space.conformal_linear_map` * `analysis.normed_space.linear_isometry` * `analysis.normed_space.star.basic` * `analysis.seminorm` * `category_theory.abelian.opposite` * `category_theory.abelian.subobject` * `category_theory.adjunction.adjoint_functor_theorems` * `category_theory.concrete_category.unbundled_hom` * `category_theory.generator` * `category_theory.limits.constructions.weakly_initial` * `category_theory.limits.preserves.functor_category` * `category_theory.limits.presheaf` * `category_theory.limits.shapes.wide_equalizers` * `category_theory.linear.yoneda` * `category_theory.preadditive.generator` * `category_theory.preadditive.yoneda.basic` * `category_theory.subobject.comma` * `data.real.hyperreal` * `data.string.defs` * `linear_algebra.affine_space.basis` * `linear_algebra.finite_dimensional` * `linear_algebra.matrix.absolute_value` * `linear_algebra.matrix.circulant` * `linear_algebra.matrix.nonsingular_inverse` * `linear_algebra.vandermonde` * `topology.extremally_disconnected` * `topology.instances.matrix` --- src/algebra/category/Module/tannaka.lean | 3 +++ src/algebra/homology/homotopy_category.lean | 3 +++ src/algebra/homology/short_exact/preadditive.lean | 3 +++ src/algebraic_topology/dold_kan/equivalence_additive.lean | 5 ++++- src/algebraic_topology/dold_kan/functor_gamma.lean | 3 +++ src/algebraic_topology/dold_kan/gamma_comp_n.lean | 5 ++++- src/algebraic_topology/dold_kan/homotopy_equivalence.lean | 3 +++ src/algebraic_topology/dold_kan/n_comp_gamma.lean | 5 ++++- src/algebraic_topology/dold_kan/normalized.lean | 3 +++ src/analysis/box_integral/partition/basic.lean | 3 +++ src/analysis/box_integral/partition/split.lean | 3 +++ src/analysis/convex/caratheodory.lean | 3 +++ src/analysis/convex/combination.lean | 3 +++ src/analysis/convex/independent.lean | 3 +++ src/analysis/convex/jensen.lean | 3 +++ src/analysis/convex/join.lean | 3 +++ src/analysis/convex/normed.lean | 3 +++ src/analysis/convex/simplicial_complex/basic.lean | 3 +++ src/analysis/convex/topology.lean | 3 +++ src/analysis/locally_convex/bounded.lean | 3 +++ src/analysis/normed_space/conformal_linear_map.lean | 3 +++ src/analysis/normed_space/linear_isometry.lean | 3 +++ src/analysis/normed_space/star/basic.lean | 3 +++ src/analysis/seminorm.lean | 3 +++ src/category_theory/abelian/opposite.lean | 3 +++ src/category_theory/abelian/subobject.lean | 3 +++ src/category_theory/adjunction/adjoint_functor_theorems.lean | 3 +++ src/category_theory/concrete_category/unbundled_hom.lean | 3 +++ src/category_theory/generator.lean | 3 +++ src/category_theory/limits/constructions/weakly_initial.lean | 3 +++ src/category_theory/limits/preserves/functor_category.lean | 3 +++ src/category_theory/limits/presheaf.lean | 3 +++ src/category_theory/limits/shapes/wide_equalizers.lean | 3 +++ src/category_theory/linear/yoneda.lean | 3 +++ src/category_theory/preadditive/generator.lean | 3 +++ src/category_theory/preadditive/yoneda/basic.lean | 3 +++ src/category_theory/subobject/comma.lean | 3 +++ src/data/real/hyperreal.lean | 3 +++ src/data/string/defs.lean | 3 +++ src/linear_algebra/affine_space/basis.lean | 3 +++ src/linear_algebra/finite_dimensional.lean | 3 +++ src/linear_algebra/matrix/absolute_value.lean | 3 +++ src/linear_algebra/matrix/circulant.lean | 3 +++ src/linear_algebra/matrix/nonsingular_inverse.lean | 3 +++ src/linear_algebra/vandermonde.lean | 3 +++ src/topology/extremally_disconnected.lean | 3 +++ src/topology/instances/matrix.lean | 3 +++ 47 files changed, 144 insertions(+), 3 deletions(-) diff --git a/src/algebra/category/Module/tannaka.lean b/src/algebra/category/Module/tannaka.lean index 4932673af4e7b..03dec3c757f8d 100644 --- a/src/algebra/category/Module/tannaka.lean +++ b/src/algebra/category/Module/tannaka.lean @@ -9,6 +9,9 @@ import linear_algebra.span /-! # Tannaka duality for rings +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A ring `R` is equivalent to the endomorphisms of the additive forgetful functor `Module R ⥤ AddCommGroup`. diff --git a/src/algebra/homology/homotopy_category.lean b/src/algebra/homology/homotopy_category.lean index d9c581898935e..08eed788b3823 100644 --- a/src/algebra/homology/homotopy_category.lean +++ b/src/algebra/homology/homotopy_category.lean @@ -9,6 +9,9 @@ import category_theory.quotient /-! # The homotopy category +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + `homotopy_category V c` gives the category of chain complexes of shape `c` in `V`, with chain maps identified when they are homotopic. -/ diff --git a/src/algebra/homology/short_exact/preadditive.lean b/src/algebra/homology/short_exact/preadditive.lean index aea22167cbedc..899a730744bd9 100644 --- a/src/algebra/homology/short_exact/preadditive.lean +++ b/src/algebra/homology/short_exact/preadditive.lean @@ -9,6 +9,9 @@ import category_theory.preadditive.additive_functor /-! # Short exact sequences, and splittings. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + `short_exact f g` is the proposition that `0 ⟶ A -f⟶ B -g⟶ C ⟶ 0` is an exact sequence. We define when a short exact sequence is left-split, right-split, and split. diff --git a/src/algebraic_topology/dold_kan/equivalence_additive.lean b/src/algebraic_topology/dold_kan/equivalence_additive.lean index 166b18d6bce1d..886aefc7db66b 100644 --- a/src/algebraic_topology/dold_kan/equivalence_additive.lean +++ b/src/algebraic_topology/dold_kan/equivalence_additive.lean @@ -6,7 +6,10 @@ Authors: Joël Riou import algebraic_topology.dold_kan.n_comp_gamma -/-! The Dold-Kan equivalence for additive categories. +/-! +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The Dold-Kan equivalence for additive categories. This file defines `preadditive.dold_kan.equivalence` which is the equivalence of categories `karoubi (simplicial_object C) ≌ karoubi (chain_complex C ℕ)`. diff --git a/src/algebraic_topology/dold_kan/functor_gamma.lean b/src/algebraic_topology/dold_kan/functor_gamma.lean index b35bc252de1e6..9a5675b665216 100644 --- a/src/algebraic_topology/dold_kan/functor_gamma.lean +++ b/src/algebraic_topology/dold_kan/functor_gamma.lean @@ -10,6 +10,9 @@ import algebraic_topology.dold_kan.split_simplicial_object # Construction of the inverse functor of the Dold-Kan equivalence +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we construct the functor `Γ₀ : chain_complex C ℕ ⥤ simplicial_object C` which shall be the inverse functor of the Dold-Kan equivalence in the case of abelian categories, diff --git a/src/algebraic_topology/dold_kan/gamma_comp_n.lean b/src/algebraic_topology/dold_kan/gamma_comp_n.lean index 8d4a813f03d44..b3d66c5c2cfcc 100644 --- a/src/algebraic_topology/dold_kan/gamma_comp_n.lean +++ b/src/algebraic_topology/dold_kan/gamma_comp_n.lean @@ -7,7 +7,10 @@ Authors: Joël Riou import algebraic_topology.dold_kan.functor_gamma import category_theory.idempotents.homological_complex -/-! The counit isomorphism of the Dold-Kan equivalence +/-! +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The counit isomorphism of the Dold-Kan equivalence The purpose of this file is to construct natural isomorphisms `N₁Γ₀ : Γ₀ ⋙ N₁ ≅ to_karoubi (chain_complex C ℕ)` diff --git a/src/algebraic_topology/dold_kan/homotopy_equivalence.lean b/src/algebraic_topology/dold_kan/homotopy_equivalence.lean index da71ea0ff7381..8db3a58d5be9b 100644 --- a/src/algebraic_topology/dold_kan/homotopy_equivalence.lean +++ b/src/algebraic_topology/dold_kan/homotopy_equivalence.lean @@ -10,6 +10,9 @@ import algebraic_topology.dold_kan.normalized # The normalized Moore complex and the alternating face map complex are homotopy equivalent +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, when the category `A` is abelian, we obtain the homotopy equivalence `homotopy_equiv_normalized_Moore_complex_alternating_face_map_complex` between the normalized Moore complex and the alternating face map complex of a simplicial object in `A`. diff --git a/src/algebraic_topology/dold_kan/n_comp_gamma.lean b/src/algebraic_topology/dold_kan/n_comp_gamma.lean index 415233b39b072..52bdb21046dc0 100644 --- a/src/algebraic_topology/dold_kan/n_comp_gamma.lean +++ b/src/algebraic_topology/dold_kan/n_comp_gamma.lean @@ -7,7 +7,10 @@ Authors: Joël Riou import algebraic_topology.dold_kan.gamma_comp_n import algebraic_topology.dold_kan.n_reflects_iso -/-! The unit isomorphism of the Dold-Kan equivalence +/-! +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The unit isomorphism of the Dold-Kan equivalence In order to construct the unit isomorphism of the Dold-Kan equivalence, we first construct natural transformations diff --git a/src/algebraic_topology/dold_kan/normalized.lean b/src/algebraic_topology/dold_kan/normalized.lean index 8a70211b6bf52..c182615e41b7c 100644 --- a/src/algebraic_topology/dold_kan/normalized.lean +++ b/src/algebraic_topology/dold_kan/normalized.lean @@ -10,6 +10,9 @@ import algebraic_topology.dold_kan.functor_n # Comparison with the normalized Moore complex functor +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + TODO (@joelriou) continue adding the various files referenced below In this file, we show that when the category `A` is abelian, diff --git a/src/analysis/box_integral/partition/basic.lean b/src/analysis/box_integral/partition/basic.lean index 81704d3638187..ce849353626fc 100644 --- a/src/analysis/box_integral/partition/basic.lean +++ b/src/analysis/box_integral/partition/basic.lean @@ -9,6 +9,9 @@ import analysis.box_integral.box.basic /-! # Partitions of rectangular boxes in `ℝⁿ` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define (pre)partitions of rectangular boxes in `ℝⁿ`. A partition of a box `I` in `ℝⁿ` (see `box_integral.prepartition` and `box_integral.prepartition.is_partition`) is a finite set of pairwise disjoint boxes such that their union is exactly `I`. We use `boxes : finset (box ι)` to diff --git a/src/analysis/box_integral/partition/split.lean b/src/analysis/box_integral/partition/split.lean index 1e3e0ad4a63a0..45dbc69989c9e 100644 --- a/src/analysis/box_integral/partition/split.lean +++ b/src/analysis/box_integral/partition/split.lean @@ -8,6 +8,9 @@ import analysis.box_integral.partition.basic /-! # Split a box along one or more hyperplanes +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions A hyperplane `{x : ι → ℝ | x i = a}` splits a rectangular box `I : box_integral.box ι` into two diff --git a/src/analysis/convex/caratheodory.lean b/src/analysis/convex/caratheodory.lean index 1ffbb069d37cd..850b28641e0d3 100644 --- a/src/analysis/convex/caratheodory.lean +++ b/src/analysis/convex/caratheodory.lean @@ -10,6 +10,9 @@ import tactic.field_simp /-! # Carathéodory's convexity theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Convex hull can be regarded as a refinement of affine span. Both are closure operators but whereas convex hull takes values in the lattice of convex subsets, affine span takes values in the much coarser sublattice of affine subspaces. diff --git a/src/analysis/convex/combination.lean b/src/analysis/convex/combination.lean index ee1a808317e50..8cc8ae375c102 100644 --- a/src/analysis/convex/combination.lean +++ b/src/analysis/convex/combination.lean @@ -10,6 +10,9 @@ import linear_algebra.affine_space.basis /-! # Convex combinations +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines convex combinations of points in a vector space. ## Main declarations diff --git a/src/analysis/convex/independent.lean b/src/analysis/convex/independent.lean index 42e5193fbc833..146ec06b43de8 100644 --- a/src/analysis/convex/independent.lean +++ b/src/analysis/convex/independent.lean @@ -9,6 +9,9 @@ import analysis.convex.extreme /-! # Convex independence +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines convex independent families of points. Convex independence is closely related to affine independence. In both cases, no point can be diff --git a/src/analysis/convex/jensen.lean b/src/analysis/convex/jensen.lean index 6dad2a1e76a82..e65f7726f8ea6 100644 --- a/src/analysis/convex/jensen.lean +++ b/src/analysis/convex/jensen.lean @@ -9,6 +9,9 @@ import analysis.convex.function /-! # Jensen's inequality and maximum principle for convex functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we prove the finite Jensen inequality and the finite maximum principle for convex functions. The integral versions are to be found in `analysis.convex.integral`. diff --git a/src/analysis/convex/join.lean b/src/analysis/convex/join.lean index 1080ba9973642..1a0ec117a51c7 100644 --- a/src/analysis/convex/join.lean +++ b/src/analysis/convex/join.lean @@ -8,6 +8,9 @@ import analysis.convex.combination /-! # Convex join +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the convex join of two sets. The convex join of `s` and `t` is the union of the segments with one end in `s` and the other in `t`. This is notably a useful gadget to deal with convex hulls of finite sets. diff --git a/src/analysis/convex/normed.lean b/src/analysis/convex/normed.lean index b5ac0d553abd5..91399a8f31073 100644 --- a/src/analysis/convex/normed.lean +++ b/src/analysis/convex/normed.lean @@ -11,6 +11,9 @@ import analysis.normed_space.ray /-! # Topological and metric properties of convex sets in normed spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We prove the following facts: * `convex_on_norm`, `convex_on_dist` : norm and distance to a fixed point is convex on any convex diff --git a/src/analysis/convex/simplicial_complex/basic.lean b/src/analysis/convex/simplicial_complex/basic.lean index 44e752b7ddf09..f29d4b3381658 100644 --- a/src/analysis/convex/simplicial_complex/basic.lean +++ b/src/analysis/convex/simplicial_complex/basic.lean @@ -9,6 +9,9 @@ import linear_algebra.affine_space.independent /-! # Simplicial complexes +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we define simplicial complexes in `𝕜`-modules. A simplicial complex is a collection of simplices closed by inclusion (of vertices) and intersection (of underlying sets). diff --git a/src/analysis/convex/topology.lean b/src/analysis/convex/topology.lean index cc9d7665e1d2a..c8248ad88ad89 100644 --- a/src/analysis/convex/topology.lean +++ b/src/analysis/convex/topology.lean @@ -12,6 +12,9 @@ import topology.algebra.module.basic /-! # Topological properties of convex sets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We prove the following facts: * `convex.interior` : interior of a convex set is convex; diff --git a/src/analysis/locally_convex/bounded.lean b/src/analysis/locally_convex/bounded.lean index 795f33df4e59a..d44d4d45c5eb3 100644 --- a/src/analysis/locally_convex/bounded.lean +++ b/src/analysis/locally_convex/bounded.lean @@ -13,6 +13,9 @@ import topology.uniform_space.cauchy /-! # Von Neumann Boundedness +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines natural or von Neumann bounded sets and proves elementary properties. ## Main declarations diff --git a/src/analysis/normed_space/conformal_linear_map.lean b/src/analysis/normed_space/conformal_linear_map.lean index 9f66f89da2471..a711a152372cb 100644 --- a/src/analysis/normed_space/conformal_linear_map.lean +++ b/src/analysis/normed_space/conformal_linear_map.lean @@ -9,6 +9,9 @@ import analysis.normed_space.linear_isometry /-! # Conformal Linear Maps +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A continuous linear map between `R`-normed spaces `X` and `Y` `is_conformal_map` if it is a nonzero multiple of a linear isometry. diff --git a/src/analysis/normed_space/linear_isometry.lean b/src/analysis/normed_space/linear_isometry.lean index b4088ddb2268a..7e87a92a9ae37 100644 --- a/src/analysis/normed_space/linear_isometry.lean +++ b/src/analysis/normed_space/linear_isometry.lean @@ -10,6 +10,9 @@ import linear_algebra.basis /-! # (Semi-)linear isometries +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define `linear_isometry σ₁₂ E E₂` (notation: `E →ₛₗᵢ[σ₁₂] E₂`) to be a semilinear isometric embedding of `E` into `E₂` and `linear_isometry_equiv` (notation: `E ≃ₛₗᵢ[σ₁₂] E₂`) to be a semilinear isometric equivalence between `E` and `E₂`. The notation for the associated purely diff --git a/src/analysis/normed_space/star/basic.lean b/src/analysis/normed_space/star/basic.lean index fe044ee76b36c..b19ac63146511 100644 --- a/src/analysis/normed_space/star/basic.lean +++ b/src/analysis/normed_space/star/basic.lean @@ -14,6 +14,9 @@ import topology.algebra.star_subalgebra /-! # Normed star rings and algebras +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A normed star group is a normed group with a compatible `star` which is isometric. A C⋆-ring is a normed star group that is also a ring and that verifies the stronger diff --git a/src/analysis/seminorm.lean b/src/analysis/seminorm.lean index 517b0787f6322..b55e5f5b27a5a 100644 --- a/src/analysis/seminorm.lean +++ b/src/analysis/seminorm.lean @@ -11,6 +11,9 @@ import analysis.normed.group.add_torsor /-! # Seminorms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines seminorms. A seminorm is a function to the reals which is positive-semidefinite, absolutely homogeneous, and diff --git a/src/category_theory/abelian/opposite.lean b/src/category_theory/abelian/opposite.lean index 6b8fe7955ef6c..7909054b7506f 100644 --- a/src/category_theory/abelian/opposite.lean +++ b/src/category_theory/abelian/opposite.lean @@ -9,6 +9,9 @@ import category_theory.limits.opposites /-! # The opposite of an abelian category is abelian. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ noncomputable theory diff --git a/src/category_theory/abelian/subobject.lean b/src/category_theory/abelian/subobject.lean index 02997c657fea5..202ac19459a56 100644 --- a/src/category_theory/abelian/subobject.lean +++ b/src/category_theory/abelian/subobject.lean @@ -9,6 +9,9 @@ import category_theory.abelian.basic /-! # Equivalence between subobjects and quotients in an abelian category +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + -/ open category_theory category_theory.limits opposite diff --git a/src/category_theory/adjunction/adjoint_functor_theorems.lean b/src/category_theory/adjunction/adjoint_functor_theorems.lean index b13c7016fb27c..9b99b59f433ab 100644 --- a/src/category_theory/adjunction/adjoint_functor_theorems.lean +++ b/src/category_theory/adjunction/adjoint_functor_theorems.lean @@ -12,6 +12,9 @@ import category_theory.subobject.comma /-! # Adjoint functor theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves the (general) adjoint functor theorem, in the form: * If `G : D ⥤ C` preserves limits and `D` has limits, and satisfies the solution set condition, then it has a left adjoint: `is_right_adjoint_of_preserves_limits_of_solution_set_condition`. diff --git a/src/category_theory/concrete_category/unbundled_hom.lean b/src/category_theory/concrete_category/unbundled_hom.lean index ba9dab6848131..22237c3b072b5 100644 --- a/src/category_theory/concrete_category/unbundled_hom.lean +++ b/src/category_theory/concrete_category/unbundled_hom.lean @@ -8,6 +8,9 @@ import category_theory.concrete_category.bundled_hom /-! # Category instances for structures that use unbundled homs +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file provides basic infrastructure to define concrete categories using unbundled homs (see `class unbundled_hom`), and define forgetful functors between them (see diff --git a/src/category_theory/generator.lean b/src/category_theory/generator.lean index 1ba6e9b0a90c1..5f79ceb63f136 100644 --- a/src/category_theory/generator.lean +++ b/src/category_theory/generator.lean @@ -14,6 +14,9 @@ import data.set.opposite /-! # Separating and detecting sets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + There are several non-equivalent notions of a generator of a category. Here, we consider two of them: diff --git a/src/category_theory/limits/constructions/weakly_initial.lean b/src/category_theory/limits/constructions/weakly_initial.lean index 2995d575ff221..54e737517a14c 100644 --- a/src/category_theory/limits/constructions/weakly_initial.lean +++ b/src/category_theory/limits/constructions/weakly_initial.lean @@ -10,6 +10,9 @@ import category_theory.limits.shapes.terminal /-! # Constructions related to weakly initial objects +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file gives constructions related to weakly initial objects, namely: * If a category has small products and a small weakly initial set of objects, then it has a weakly initial object. diff --git a/src/category_theory/limits/preserves/functor_category.lean b/src/category_theory/limits/preserves/functor_category.lean index a267cd407e0df..0770cdaa956a1 100644 --- a/src/category_theory/limits/preserves/functor_category.lean +++ b/src/category_theory/limits/preserves/functor_category.lean @@ -11,6 +11,9 @@ import category_theory.limits.presheaf /-! # Preservation of (co)limits in the functor category +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + * Show that if `X ⨯ -` preserves colimits in `D` for any `X : D`, then the product functor `F ⨯ -` for `F : C ⥤ D` preserves colimits. diff --git a/src/category_theory/limits/presheaf.lean b/src/category_theory/limits/presheaf.lean index b18e8cf46220c..166e034b49100 100644 --- a/src/category_theory/limits/presheaf.lean +++ b/src/category_theory/limits/presheaf.lean @@ -14,6 +14,9 @@ import category_theory.limits.types /-! # Colimit of representables +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file constructs an adjunction `yoneda_adjunction` between `(Cᵒᵖ ⥤ Type u)` and `ℰ` given a functor `A : C ⥤ ℰ`, where the right adjoint sends `(E : ℰ)` to `c ↦ (A.obj c ⟶ E)` (provided `ℰ` has colimits). diff --git a/src/category_theory/limits/shapes/wide_equalizers.lean b/src/category_theory/limits/shapes/wide_equalizers.lean index 4befef2d48735..1ffed64a6845c 100644 --- a/src/category_theory/limits/shapes/wide_equalizers.lean +++ b/src/category_theory/limits/shapes/wide_equalizers.lean @@ -9,6 +9,9 @@ import category_theory.limits.shapes.equalizers /-! # Wide equalizers and wide coequalizers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines wide (co)equalizers as special cases of (co)limits. A wide equalizer for the family of morphisms `X ⟶ Y` indexed by `J` is the categorical diff --git a/src/category_theory/linear/yoneda.lean b/src/category_theory/linear/yoneda.lean index 159a6a28c6b28..c888432d38571 100644 --- a/src/category_theory/linear/yoneda.lean +++ b/src/category_theory/linear/yoneda.lean @@ -10,6 +10,9 @@ import category_theory.preadditive.yoneda.basic /-! # The Yoneda embedding for `R`-linear categories +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The Yoneda embedding for `R`-linear categories `C`, sends an object `X : C` to the `Module R`-valued presheaf on `C`, with value on `Y : Cᵒᵖ` given by `Module.of R (unop Y ⟶ X)`. diff --git a/src/category_theory/preadditive/generator.lean b/src/category_theory/preadditive/generator.lean index ddae1e7b7633e..c096257eed802 100644 --- a/src/category_theory/preadditive/generator.lean +++ b/src/category_theory/preadditive/generator.lean @@ -9,6 +9,9 @@ import category_theory.preadditive.yoneda.basic /-! # Separators in preadditive categories +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains characterizations of separating sets and objects that are valid in all preadditive categories. diff --git a/src/category_theory/preadditive/yoneda/basic.lean b/src/category_theory/preadditive/yoneda/basic.lean index 792a908f13313..64cd6f6302f84 100644 --- a/src/category_theory/preadditive/yoneda/basic.lean +++ b/src/category_theory/preadditive/yoneda/basic.lean @@ -11,6 +11,9 @@ import algebra.category.Group.preadditive /-! # The Yoneda embedding for preadditive categories +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The Yoneda embedding for preadditive categories sends an object `Y` to the presheaf sending an object `X` to the group of morphisms `X ⟶ Y`. At each point, we get an additional `End Y`-module structure. diff --git a/src/category_theory/subobject/comma.lean b/src/category_theory/subobject/comma.lean index 76222c511100b..3c74b8cf0034c 100644 --- a/src/category_theory/subobject/comma.lean +++ b/src/category_theory/subobject/comma.lean @@ -10,6 +10,9 @@ import category_theory.limits.shapes.finite_limits /-! # Subobjects in the category of structured arrows +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We compute the subobjects of an object `A` in the category `structured_arrow S T` for `T : C ⥤ D` and `S : D` as a subtype of the subobjects of `A.right`. We deduce that `structured_arrow S T` is well-powered if `C` is. diff --git a/src/data/real/hyperreal.lean b/src/data/real/hyperreal.lean index 054a124f0077a..fc7dcc328ccfb 100644 --- a/src/data/real/hyperreal.lean +++ b/src/data/real/hyperreal.lean @@ -8,6 +8,9 @@ import analysis.specific_limits.basic /-! # Construction of the hyperreal numbers as an ultraproduct of real sequences. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ open filter filter.germ diff --git a/src/data/string/defs.lean b/src/data/string/defs.lean index 5dcca2d98d967..ea5351bf282e7 100644 --- a/src/data/string/defs.lean +++ b/src/data/string/defs.lean @@ -8,6 +8,9 @@ import data.list.defs /-! # Definitions for `string` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines a bunch of functions for the `string` datatype. -/ diff --git a/src/linear_algebra/affine_space/basis.lean b/src/linear_algebra/affine_space/basis.lean index 42e4192d7ea65..2dbba32d15418 100644 --- a/src/linear_algebra/affine_space/basis.lean +++ b/src/linear_algebra/affine_space/basis.lean @@ -9,6 +9,9 @@ import linear_algebra.basis /-! # Affine bases and barycentric coordinates +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Suppose `P` is an affine space modelled on the module `V` over the ring `k`, and `p : ι → P` is an affine-independent family of points spanning `P`. Given this data, each point `q : P` may be written uniquely as an affine combination: `q = w₀ p₀ + w₁ p₁ + ⋯` for some (finitely-supported) weights diff --git a/src/linear_algebra/finite_dimensional.lean b/src/linear_algebra/finite_dimensional.lean index efbac63192c3d..a85cdf9c428b7 100644 --- a/src/linear_algebra/finite_dimensional.lean +++ b/src/linear_algebra/finite_dimensional.lean @@ -11,6 +11,9 @@ import tactic.interval_cases /-! # Finite dimensional vector spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Definition and basic properties of finite dimensional vector spaces, of their dimensions, and of linear maps on such spaces. diff --git a/src/linear_algebra/matrix/absolute_value.lean b/src/linear_algebra/matrix/absolute_value.lean index 3ebb90dbc68eb..38a7a0881fffe 100644 --- a/src/linear_algebra/matrix/absolute_value.lean +++ b/src/linear_algebra/matrix/absolute_value.lean @@ -9,6 +9,9 @@ import linear_algebra.matrix.determinant /-! # Absolute values and matrices +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves some bounds on matrices involving absolute values. ## Main results diff --git a/src/linear_algebra/matrix/circulant.lean b/src/linear_algebra/matrix/circulant.lean index 6b603bff25d5b..8ee4942ac5a66 100644 --- a/src/linear_algebra/matrix/circulant.lean +++ b/src/linear_algebra/matrix/circulant.lean @@ -8,6 +8,9 @@ import linear_algebra.matrix.symmetric /-! # Circulant matrices +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains the definition and basic results about circulant matrices. Given a vector `v : n → α` indexed by a type that is endowed with subtraction, `matrix.circulant v` is the matrix whose `(i, j)`th entry is `v (i - j)`. diff --git a/src/linear_algebra/matrix/nonsingular_inverse.lean b/src/linear_algebra/matrix/nonsingular_inverse.lean index c730f21c63493..66b7ddd6cdcaa 100644 --- a/src/linear_algebra/matrix/nonsingular_inverse.lean +++ b/src/linear_algebra/matrix/nonsingular_inverse.lean @@ -8,6 +8,9 @@ import linear_algebra.matrix.adjugate /-! # Nonsingular inverses +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we define an inverse for square matrices of invertible determinant. For matrices that are not square or not of full rank, there is a more general notion of diff --git a/src/linear_algebra/vandermonde.lean b/src/linear_algebra/vandermonde.lean index 4b473cf4a6b61..ff3ce3eb48755 100644 --- a/src/linear_algebra/vandermonde.lean +++ b/src/linear_algebra/vandermonde.lean @@ -11,6 +11,9 @@ import linear_algebra.matrix.nondegenerate /-! # Vandermonde matrix +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the `vandermonde` matrix and gives its determinant. ## Main definitions diff --git a/src/topology/extremally_disconnected.lean b/src/topology/extremally_disconnected.lean index e09313a7f6680..c5796de9c98d7 100644 --- a/src/topology/extremally_disconnected.lean +++ b/src/topology/extremally_disconnected.lean @@ -8,6 +8,9 @@ import topology.stone_cech /-! # Extremally disconnected spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + An extremally disconnected topological space is a space in which the closure of every open set is open. Such spaces are also called Stonean spaces. They are the projective objects in the category of compact Hausdorff spaces. diff --git a/src/topology/instances/matrix.lean b/src/topology/instances/matrix.lean index dd801b2f40563..ad02ee006beb2 100644 --- a/src/topology/instances/matrix.lean +++ b/src/topology/instances/matrix.lean @@ -12,6 +12,9 @@ import linear_algebra.matrix.trace /-! # Topological properties of matrices +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file is a place to collect topological results about matrices. ## Main definitions: From 04cdee31e196e30f507e8e9eb2d06e02c9ff6310 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Thu, 27 Apr 2023 08:43:40 +0000 Subject: [PATCH 0886/1291] feat(linear_algebra/basis): add basis.restrict_scalars (#18814) Co-authored-by: Eric Wieser --- src/linear_algebra/basis.lean | 49 +++++++++++++++++++++++++++++++++++ 1 file changed, 49 insertions(+) diff --git a/src/linear_algebra/basis.lean b/src/linear_algebra/basis.lean index 6bd54cd42e48b..c560d1cfc33ec 100644 --- a/src/linear_algebra/basis.lean +++ b/src/linear_algebra/basis.lean @@ -1424,3 +1424,52 @@ let ⟨q, hq⟩ := p.exists_is_compl in nonempty.intro $ (prod_equiv_of_is_compl q p hq.symm) end division_ring + +section restrict_scalars + +variables {S : Type*} [comm_ring R] [ring S] [nontrivial S] [add_comm_group M] +variables [algebra R S] [module S M] [module R M] +variables [is_scalar_tower R S M] [no_zero_smul_divisors R S] (b : basis ι S M) +variables (R) + +open submodule + +/-- Let `b` be a `S`-basis of `M`. Let `R` be a comm_ring such that `algebra R S` with no zero +smul divisors, then the submodule of `M` spanned by `b` over `R` admits `b` as a `R`-basis. -/ +noncomputable def basis.restrict_scalars : basis ι R (span R (set.range b)) := +basis.span (b.linear_independent.restrict_scalars (smul_left_injective R one_ne_zero)) + +@[simp] +lemma basis.restrict_scalars_apply (i : ι) : (b.restrict_scalars R i : M) = b i := +by simp only [basis.restrict_scalars, basis.span_apply] + +@[simp] +lemma basis.restrict_scalars_repr_apply (m : span R (set.range b)) (i : ι) : + algebra_map R S ((b.restrict_scalars R).repr m i) = b.repr m i := +begin + suffices : finsupp.map_range.linear_map (algebra.linear_map R S) ∘ₗ + (b.restrict_scalars R).repr.to_linear_map + = ((b.repr : M →ₗ[S] (ι →₀ S)).restrict_scalars R).dom_restrict _, + { exact finsupp.congr_fun (linear_map.congr_fun this m) i, }, + refine basis.ext (b.restrict_scalars R) (λ _, _), + simp only [linear_map.coe_comp, linear_equiv.coe_to_linear_map, function.comp_app, map_one, + basis.repr_self, finsupp.map_range.linear_map_apply, finsupp.map_range_single, + algebra.linear_map_apply, linear_map.dom_restrict_apply, linear_equiv.coe_coe, + basis.restrict_scalars_apply, linear_map.coe_restrict_scalars_eq_coe], +end + +/-- Let `b` be a `S`-basis of `M`. Then `m : M` lies in the `R`-module spanned by `b` iff all the +coordinates of `m` on the basis `b` are in `R` (see `basis.mem_span` for the case `R = S`). -/ +lemma basis.mem_span_iff_repr_mem (m : M) : + m ∈ span R (set.range b) ↔ ∀ i, b.repr m i ∈ set.range (algebra_map R S) := +begin + refine ⟨λ hm i, ⟨(b.restrict_scalars R).repr ⟨m, hm⟩ i, + (b.restrict_scalars_repr_apply R ⟨m, hm⟩ i)⟩, λ h, _⟩, + rw [← b.total_repr m, finsupp.total_apply S _], + refine sum_mem (λ i _, _), + obtain ⟨_, h⟩ := h i, + simp_rw [← h, algebra_map_smul], + exact smul_mem _ _ (subset_span (set.mem_range_self i)), +end + +end restrict_scalars From fa78268d4d77cb2b2fbc89f0527e2e7807763780 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 27 Apr 2023 11:18:50 +0000 Subject: [PATCH 0887/1291] chore(ring_theory/finiteness): generalize `module.finite.trans` (#18880) This was stated for algebras but holds more generally for modules. This now can be used to prove `finite_dimensional.trans`. A few downstream proofs were providing the instances arguments explicitly and consequently broke. Passing those instances via `haveI` fixes those proofs and makes them less fragile. --- src/field_theory/tower.lean | 3 +-- src/number_theory/cyclotomic/basic.lean | 8 +++++--- src/ring_theory/finiteness.lean | 23 ++++++++++++----------- 3 files changed, 18 insertions(+), 16 deletions(-) diff --git a/src/field_theory/tower.lean b/src/field_theory/tower.lean index 8c67f3c076912..66c8e28cc56dc 100644 --- a/src/field_theory/tower.lean +++ b/src/field_theory/tower.lean @@ -65,8 +65,7 @@ namespace finite_dimensional open is_noetherian theorem trans [finite_dimensional F K] [finite_dimensional K A] : finite_dimensional F A := -let b := basis.of_vector_space F K, c := basis.of_vector_space K A in -of_fintype_basis $ b.smul c +module.finite.trans K A /-- In a tower of field extensions `L / K / F`, if `L / F` is finite, so is `K / F`. diff --git a/src/number_theory/cyclotomic/basic.lean b/src/number_theory/cyclotomic/basic.lean index 966eba91a1f36..2d936346d0181 100644 --- a/src/number_theory/cyclotomic/basic.lean +++ b/src/number_theory/cyclotomic/basic.lean @@ -339,9 +339,11 @@ end lemma number_field [h : number_field K] [_root_.finite S] [is_cyclotomic_extension S K L] : number_field L := { to_char_zero := char_zero_of_injective_algebra_map (algebra_map K L).injective, - to_finite_dimensional := @module.finite.trans _ K L _ _ _ _ - (@algebra_rat L _ (char_zero_of_injective_algebra_map (algebra_map K L).injective)) _ _ - h.to_finite_dimensional (finite S K L) } + to_finite_dimensional := begin + haveI := char_zero_of_injective_algebra_map (algebra_map K L).injective, + haveI := finite S K L, + exact module.finite.trans K _ + end } localized "attribute [instance] is_cyclotomic_extension.number_field" in cyclotomic diff --git a/src/ring_theory/finiteness.lean b/src/ring_theory/finiteness.lean index 1b22d2e24e4f0..a337c43640ef1 100644 --- a/src/ring_theory/finiteness.lean +++ b/src/ring_theory/finiteness.lean @@ -522,13 +522,13 @@ of_surjective (e : M →ₗ[R] N) e.surjective section algebra -lemma trans {R : Type*} (A B : Type*) [comm_semiring R] [comm_semiring A] [algebra R A] - [semiring B] [algebra R B] [algebra A B] [is_scalar_tower R A B] : - ∀ [finite R A] [finite A B], finite R B +lemma trans {R : Type*} (A M : Type*) [comm_semiring R] [semiring A] [algebra R A] + [add_comm_monoid M] [module R M] [module A M] [is_scalar_tower R A M] : + ∀ [finite R A] [finite A M], finite R M | ⟨⟨s, hs⟩⟩ ⟨⟨t, ht⟩⟩ := ⟨submodule.fg_def.2 - ⟨set.image2 (•) (↑s : set A) (↑t : set B), + ⟨set.image2 (•) (↑s : set A) (↑t : set M), set.finite.image2 _ s.finite_to_set t.finite_to_set, - by rw [set.image2_smul, submodule.span_smul_of_span_eq_top hs (↑t : set B), + by rw [set.image2_smul, submodule.span_smul_of_span_eq_top hs (↑t : set M), ht, submodule.restrict_scalars_top]⟩⟩ end algebra @@ -584,14 +584,15 @@ begin end lemma comp {g : B →+* C} {f : A →+* B} (hg : g.finite) (hf : f.finite) : (g.comp f).finite := -@module.finite.trans A B C _ _ f.to_algebra _ (g.comp f).to_algebra g.to_algebra begin - fconstructor, - intros a b c, - simp only [algebra.smul_def, ring_hom.map_mul, mul_assoc], - refl + letI := f.to_algebra, + letI := g.to_algebra, + letI := (g.comp f).to_algebra, + letI : is_scalar_tower A B C := restrict_scalars.is_scalar_tower A B C, + letI : module.finite A B := hf, + letI : module.finite B C := hg, + exact module.finite.trans B C, end -hf hg lemma of_comp_finite {f : A →+* B} {g : B →+* C} (h : (g.comp f).finite) : g.finite := begin From b17e272059b7ba6d9ba850e274b08d2a2cde3ccf Mon Sep 17 00:00:00 2001 From: lukemantle Date: Thu, 27 Apr 2023 21:07:06 +0000 Subject: [PATCH 0888/1291] feat(ring_theory/polynomial): add properties of `coeff (hermite n)` (#18837) Added lemmas showing recursive relations between coefficients of `polynomial.hermite n`: - `coeff (hermite (n + 1)) 0 = -(coeff (hermite n) 1)` - `coeff (hermite (n + 1)) (k + 1) = coeff (hermite n) k - (k + 2) * (coeff (hermite n) (k + 2))` Showed some more properties of the coefficients of Hermite polynomials: - for `k>n`, the `k`th coefficient of the `n`th Hermite polynomial is zero - the `n`th Hermite polynomial is of degree `n` - the Hermite polynomials are monic Co-authored-by: Jake Levinson [levinson.jake@gmail.com](mailto:levinson.jake@gmail.com) Co-authored-by: lukemantle <62187700+lukemantle@users.noreply.github.com> Co-authored-by: Eric Wieser --- src/ring_theory/polynomial/hermite.lean | 75 +++++++++++++++++++++++++ 1 file changed, 75 insertions(+) diff --git a/src/ring_theory/polynomial/hermite.lean b/src/ring_theory/polynomial/hermite.lean index a1a5d7f677c91..92699c9b6594f 100644 --- a/src/ring_theory/polynomial/hermite.lean +++ b/src/ring_theory/polynomial/hermite.lean @@ -5,6 +5,7 @@ Authors: Luke Mantle -/ import data.polynomial.derivative +import data.nat.parity /-! # Hermite polynomials @@ -16,6 +17,12 @@ This file defines `polynomial.hermite n`, the nth probabilist's Hermite polynomi * `polynomial.hermite n`: the `n`th probabilist's Hermite polynomial, defined recursively as a `polynomial ℤ` +## Results + +* `polynomial.coeff_hermite_of_odd_add`: for `n`,`k` where `n+k` is odd, `(hermite n).coeff k` is + zero. +* `polynomial.monic_hermite`: for all `n`, `hermite n` is monic. + ## References * [Hermite Polynomials](https://en.wikipedia.org/wiki/Hermite_polynomials) @@ -50,4 +57,72 @@ begin simp only [map_one, mul_one, derivative_one, sub_zero] end +/-! ### Lemmas about `polynomial.coeff` -/ + +section coeff + +lemma coeff_hermite_succ_zero (n : ℕ) : + coeff (hermite (n + 1)) 0 = -(coeff (hermite n) 1) := by simp [coeff_derivative] + +lemma coeff_hermite_succ_succ (n k : ℕ) : + coeff (hermite (n + 1)) (k + 1) = coeff (hermite n) k - (k + 2) * (coeff (hermite n) (k + 2)) := +begin + rw [hermite_succ, coeff_sub, coeff_X_mul, coeff_derivative, mul_comm], + norm_cast +end + +lemma coeff_hermite_of_lt {n k : ℕ} (hnk : n < k) : coeff (hermite n) k = 0 := +begin + obtain ⟨k, rfl⟩ := nat.exists_eq_add_of_lt hnk, + clear hnk, + induction n with n ih generalizing k, + { apply coeff_C }, + { have : n + k + 1 + 2 = n + (k + 2) + 1 := by ring, + rw [nat.succ_eq_add_one, coeff_hermite_succ_succ, add_right_comm, this, ih k, ih (k + 2), + mul_zero, sub_zero] } +end + +@[simp] lemma coeff_hermite_self (n : ℕ) : coeff (hermite n) n = 1 := +begin + induction n with n ih, + { apply coeff_C }, + { rw [coeff_hermite_succ_succ, ih, coeff_hermite_of_lt, mul_zero, sub_zero], + simp } +end + +@[simp] lemma degree_hermite (n : ℕ) : (hermite n).degree = n := +begin + rw degree_eq_of_le_of_coeff_ne_zero, + simp_rw [degree_le_iff_coeff_zero, with_bot.coe_lt_coe], + { intro m, + exact coeff_hermite_of_lt }, + { simp [coeff_hermite_self n] } +end + +@[simp] lemma nat_degree_hermite {n : ℕ} : (hermite n).nat_degree = n := +nat_degree_eq_of_degree_eq_some (degree_hermite n) + +@[simp] lemma leading_coeff_hermite (n : ℕ) : (hermite n).leading_coeff = 1 := +begin + rw [← coeff_nat_degree, nat_degree_hermite, coeff_hermite_self], +end + +lemma hermite_monic (n : ℕ) : (hermite n).monic := leading_coeff_hermite n + +lemma coeff_hermite_of_odd_add {n k : ℕ} (hnk : odd (n + k)) : coeff (hermite n) k = 0 := +begin + induction n with n ih generalizing k, + { rw zero_add at hnk, + exact coeff_hermite_of_lt hnk.pos }, + { cases k, + { rw nat.succ_add_eq_succ_add at hnk, + rw [coeff_hermite_succ_zero, ih hnk, neg_zero] }, + { rw [coeff_hermite_succ_succ, ih, ih, mul_zero, sub_zero], + { rwa [nat.succ_add_eq_succ_add] at hnk }, + { rw (by rw [nat.succ_add, nat.add_succ] : n.succ + k.succ = n + k + 2) at hnk, + exact (nat.odd_add.mp hnk).mpr even_two }}} +end + +end coeff + end polynomial From a5ff45a1c92c278b03b52459a620cfd9c49ebc80 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Fri, 28 Apr 2023 06:22:27 +0000 Subject: [PATCH 0889/1291] chore(category_theory/abelian): backport removal of abelian.has_finite_biproducts instance (#18740) This backports a proposed removal of the `abelian.has_finite_biproducts` global instance, instead enabling it locally in the files that need it. The reason for removing it is that it triggers the ~~dreaded~~ https://github.com/leanprover/lean4/issues/2055 during the simpNF linter in https://github.com/leanprover-community/mathlib4/pull/2769, the mathlib4 port of `category_theory.abelian.basic`. This backport verifies that we won't run into further problems downstream if we (hopefully temporarily) remove these instances in mathlib4. Co-authored-by: Scott Morrison --- src/category_theory/abelian/basic.lean | 9 +++++++-- src/category_theory/abelian/opposite.lean | 5 +++++ 2 files changed, 12 insertions(+), 2 deletions(-) diff --git a/src/category_theory/abelian/basic.lean b/src/category_theory/abelian/basic.lean index bfbd35d438ca0..c42493076843b 100644 --- a/src/category_theory/abelian/basic.lean +++ b/src/category_theory/abelian/basic.lean @@ -257,10 +257,15 @@ namespace category_theory.abelian variables {C : Type u} [category.{v} C] [abelian C] /-- An abelian category has finite biproducts. -/ -@[priority 100] -instance has_finite_biproducts : has_finite_biproducts C := +-- Porting note: this should be an instance, +-- but triggers https://github.com/leanprover/lean4/issues/2055 +-- We set it as a local instance instead. +-- @[priority 100] instance +theorem has_finite_biproducts : has_finite_biproducts C := limits.has_finite_biproducts.of_has_finite_products +local attribute [instance] has_finite_biproducts + @[priority 100] instance has_binary_biproducts : has_binary_biproducts C := limits.has_binary_biproducts_of_finite_biproducts _ diff --git a/src/category_theory/abelian/opposite.lean b/src/category_theory/abelian/opposite.lean index 7909054b7506f..e9639b75b3d69 100644 --- a/src/category_theory/abelian/opposite.lean +++ b/src/category_theory/abelian/opposite.lean @@ -25,6 +25,11 @@ variables (C : Type*) [category C] [abelian C] local attribute [instance] has_finite_limits_of_has_equalizers_and_finite_products has_finite_colimits_of_has_coequalizers_and_finite_coproducts + -- Porting note: + -- This should have been a global instance, + -- but triggers https://github.com/leanprover/lean4/issues/2055 + -- when ported to mathlib4. + abelian.has_finite_biproducts instance : abelian Cᵒᵖ := { normal_mono_of_mono := λ X Y f m, by exactI From 0b89934139d3be96f9dab477f10c20f9f93da580 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Fri, 28 Apr 2023 07:34:03 +0000 Subject: [PATCH 0890/1291] chore(mv_polynomial/funext): backport golfed proof (#18839) This proof was timing out in https://github.com/leanprover-community/mathlib4/pull/3225, so I completely rewrote it using smaller lemmas. This is the backport. Co-authored-by: Scott Morrison --- src/data/mv_polynomial/basic.lean | 16 +++++++++ src/data/mv_polynomial/funext.lean | 45 +++++++------------------- src/data/mv_polynomial/polynomial.lean | 44 +++++++++++++++++++++++++ 3 files changed, 71 insertions(+), 34 deletions(-) create mode 100644 src/data/mv_polynomial/polynomial.lean diff --git a/src/data/mv_polynomial/basic.lean b/src/data/mv_polynomial/basic.lean index 973a34c4ca479..d34f0b96e9fc0 100644 --- a/src/data/mv_polynomial/basic.lean +++ b/src/data/mv_polynomial/basic.lean @@ -901,6 +901,22 @@ begin congr' with a, simp end +@[simp] +theorem eval₂_id (p : mv_polynomial σ R) (g : σ → R) : eval₂ (ring_hom.id _) g p = eval g p := + rfl + +theorem eval_eval₂ {τ : Type*} (f : R →+* mv_polynomial τ S₁) (g : σ → mv_polynomial τ S₁) + (p : mv_polynomial σ R) (x : τ → S₁) : + eval x (eval₂ f g p) = eval₂ ((eval x).comp f) (λ s, eval x (g s)) p := +begin + apply induction_on p, + { simp, }, + { intros p q hp hq, + simp [hp, hq] }, + { intros p n hp, + simp [hp] } +end + end eval section map diff --git a/src/data/mv_polynomial/funext.lean b/src/data/mv_polynomial/funext.lean index 455d8d8ae6d8a..cf1dea8a56763 100644 --- a/src/data/mv_polynomial/funext.lean +++ b/src/data/mv_polynomial/funext.lean @@ -6,6 +6,7 @@ Authors: Johan Commelin import data.polynomial.ring_division import data.mv_polynomial.rename import ring_theory.polynomial.basic +import data.mv_polynomial.polynomial /-! ## Function extensionality for multivariate polynomials @@ -30,41 +31,17 @@ variables {R : Type*} [comm_ring R] [is_domain R] [infinite R] private lemma funext_fin {n : ℕ} {p : mv_polynomial (fin n) R} (h : ∀ x : fin n → R, eval x p = 0) : p = 0 := begin - unfreezingI { induction n with n ih generalizing R }, - { let e := (mv_polynomial.is_empty_ring_equiv R (fin 0)), - apply e.injective, - rw ring_equiv.map_zero, - convert h fin_zero_elim, - suffices : (eval₂_hom (ring_hom.id _) (is_empty.elim' fin.is_empty)) p = - (eval fin_zero_elim : mv_polynomial (fin 0) R →+* R) p, - { rw [← this], - simp only [coe_eval₂_hom, is_empty_ring_equiv_apply, - ring_equiv.trans_apply, aeval_eq_eval₂_hom], - congr }, - exact eval₂_hom_congr rfl (subsingleton.elim _ _) rfl }, - { let e := (fin_succ_equiv R n).to_ring_equiv, - apply e.injective, - simp only [ring_equiv.map_zero], - apply polynomial.funext, - intro q, + induction n with n ih, + { apply (mv_polynomial.is_empty_ring_equiv R (fin 0)).injective, + rw [ring_equiv.map_zero], + convert h _, }, + { apply (fin_succ_equiv R n).injective, + simp only [alg_equiv.map_zero], + refine polynomial.funext (λ q, _), rw [polynomial.eval_zero], - apply ih, swap, { apply_instance }, - intro x, - dsimp [e], - rw [fin_succ_equiv_apply], - calc _ = eval _ p : _ - ... = 0 : h _, - { intro i, exact fin.cases (eval x q) x i }, - apply induction_on p, - { intro r, - simp only [eval_C, polynomial.eval_C, ring_hom.coe_comp, eval₂_hom_C], }, - { intros, simp only [*, ring_hom.map_add, polynomial.eval_add] }, - { intros φ i hφ, simp only [*, eval_X, polynomial.eval_mul, ring_hom.map_mul, eval₂_hom_X'], - congr' 1, - by_cases hi : i = 0, - { subst hi, simp only [polynomial.eval_X, fin.cases_zero] }, - { rw [← fin.succ_pred i hi], simp only [eval_X, polynomial.eval_C, fin.cases_succ] } }, - { apply_instance, }, }, + apply ih (λ x, _), + calc _ = _ : eval_polynomial_eval_fin_succ_equiv p _ _ + ... = 0 : h _, } end /-- Two multivariate polynomials over an infinite integral domain are equal diff --git a/src/data/mv_polynomial/polynomial.lean b/src/data/mv_polynomial/polynomial.lean new file mode 100644 index 0000000000000..d3571c14ddbb5 --- /dev/null +++ b/src/data/mv_polynomial/polynomial.lean @@ -0,0 +1,44 @@ +/- +Copyright (c) 2023 Scott Morrison. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Scott Morrison +-/ +import data.mv_polynomial.equiv +import data.polynomial.eval + +/-! +# Some lemmas relating polynomials and multivariable polynomials. +-/ + +namespace mv_polynomial + +variables {R S : Type*} [comm_semiring R] [comm_semiring S] {σ : Type*} + +theorem polynomial_eval_eval₂ + (f : R →+* polynomial S) (g : σ → polynomial S) (p : mv_polynomial σ R) (x : S) : + polynomial.eval x (mv_polynomial.eval₂ f g p) = + mv_polynomial.eval₂ ((polynomial.eval_ring_hom x).comp f) (λ s, polynomial.eval x (g s)) p := +begin + apply mv_polynomial.induction_on p, + { simp }, + { intros p q hp hq, + simp [hp, hq], }, + { intros p n hp, + simp [hp], }, +end + +theorem eval_polynomial_eval_fin_succ_equiv {n : ℕ} + (f : mv_polynomial (fin (n + 1)) R) (q : mv_polynomial (fin n) R) (x : fin n → R) : + (eval x) (polynomial.eval q (fin_succ_equiv R n f)) = + eval (show fin (n+1) → R, from @fin.cases _ (λ _, R) (eval x q) x) f := +begin + simp only [fin_succ_equiv_apply, coe_eval₂_hom, eval_eval₂, polynomial_eval_eval₂], + have : (eval x).comp ((polynomial.eval_ring_hom q).comp (polynomial.C.comp C)) = ring_hom.id _, + { ext, simp, }, + simp only [this, eval₂_id], + congr, + funext i, + refine fin.cases (by simp) (by simp) i, +end + +end mv_polynomial From c060baa79af5ca092c54b8bf04f0f10592f59489 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 28 Apr 2023 07:34:04 +0000 Subject: [PATCH 0891/1291] chore(data/matrix/block): the block matrix of zeros is zero (#18879) --- src/data/matrix/block.lean | 4 ++++ 1 file changed, 4 insertions(+) diff --git a/src/data/matrix/block.lean b/src/data/matrix/block.lean index 0ae6e9748f3e0..34a97c6775854 100644 --- a/src/data/matrix/block.lean +++ b/src/data/matrix/block.lean @@ -202,6 +202,10 @@ begin ext i j, cases i; cases j; simp [from_blocks], end +@[simp] lemma from_blocks_zero [has_zero α] : + from_blocks (0 : matrix n l α) 0 0 (0 : matrix o m α) = 0 := +by { ext i j, rcases i; rcases j; refl } + lemma from_blocks_add [has_add α] (A : matrix n l α) (B : matrix n m α) (C : matrix o l α) (D : matrix o m α) (A' : matrix n l α) (B' : matrix n m α) (C' : matrix o l α) (D' : matrix o m α) : From 16e59248c0ebafabd5d071b1cd41743eb8698ffb Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 28 Apr 2023 09:33:34 +0000 Subject: [PATCH 0892/1291] chore(topology/continuous_function/algebra): simplify definition and proof (#18888) The definition isn't benefiting from the `coe_sort`, and the proof can save a lot of effort by staying within the subalgebra. This proof was proving annoying to port. --- src/topology/continuous_function/algebra.lean | 45 +++++++------------ .../stone_weierstrass.lean | 11 ++--- 2 files changed, 19 insertions(+), 37 deletions(-) diff --git a/src/topology/continuous_function/algebra.lean b/src/topology/continuous_function/algebra.lean index 24f792af517d2..b40b6cb18e45a 100644 --- a/src/topology/continuous_function/algebra.lean +++ b/src/topology/continuous_function/algebra.lean @@ -710,7 +710,7 @@ writing it this way avoids having to deal with casts inside the set. where the functions would be continuous functions vanishing at infinity.) -/ def set.separates_points_strongly (s : set C(α, 𝕜)) : Prop := -∀ (v : α → 𝕜) (x y : α), ∃ f : s, (f x : 𝕜) = v x ∧ f y = v y +∀ (v : α → 𝕜) (x y : α), ∃ f ∈ s, (f x : 𝕜) = v x ∧ f y = v y variables [field 𝕜] [topological_ring 𝕜] @@ -727,25 +727,15 @@ lemma subalgebra.separates_points.strongly {s : subalgebra 𝕜 C(α, 𝕜)} (h begin by_cases n : x = y, { subst n, - use ((v x) • 1 : C(α, 𝕜)), - { apply s.smul_mem, - apply s.one_mem, }, - { simp [coe_fn_coe_base'] }, }, - obtain ⟨f, ⟨f, ⟨m, rfl⟩⟩, w⟩ := h n, - replace w : f x - f y ≠ 0 := sub_ne_zero_of_ne w, + refine ⟨_, ((v x) • 1 : s).prop, mul_one _, mul_one _⟩ }, + obtain ⟨_, ⟨f, hf, rfl⟩, hxy⟩ := h n, + replace hxy : f x - f y ≠ 0 := sub_ne_zero_of_ne hxy, let a := v x, let b := v y, - let f' := ((b - a) * (f x - f y)⁻¹) • (continuous_map.C (f x) - f) + continuous_map.C a, - refine ⟨⟨f', _⟩, _, _⟩, - { simp only [f', set_like.mem_coe, subalgebra.mem_to_submodule], - -- TODO should there be a tactic for this? - -- We could add an attribute `@[subobject_mem]`, and a tactic - -- ``def subobject_mem := `[solve_by_elim with subobject_mem { max_depth := 10 }]`` - solve_by_elim - [subalgebra.add_mem, subalgebra.smul_mem, subalgebra.sub_mem, subalgebra.algebra_map_mem] - { max_depth := 6 }, }, - { simp [f', coe_fn_coe_base'], }, - { simp [f', coe_fn_coe_base', inv_mul_cancel_right₀ w], }, + let f' : s := ((b - a) * (f x - f y)⁻¹) • (algebra_map _ _ (f x) - ⟨f, hf⟩) + algebra_map _ _ a, + refine ⟨f', f'.prop, _, _⟩, + { simp [f'], }, + { simp [f', inv_mul_cancel_right₀ hxy], }, end end continuous_map @@ -753,22 +743,17 @@ end continuous_map instance continuous_map.subsingleton_subalgebra (α : Type*) [topological_space α] (R : Type*) [comm_semiring R] [topological_space R] [topological_semiring R] [subsingleton α] : subsingleton (subalgebra R C(α, R)) := -begin - fsplit, - intros s₁ s₂, - by_cases n : nonempty α, - { obtain ⟨x⟩ := n, +⟨λ s₁ s₂, begin + casesI is_empty_or_nonempty α, + { haveI : subsingleton C(α, R) := fun_like.coe_injective.subsingleton, + exact subsingleton.elim _ _ }, + { inhabit α, ext f, - have h : f = algebra_map R C(α, R) (f x), + have h : f = algebra_map R C(α, R) (f default), { ext x', simp only [mul_one, algebra.id.smul_eq_mul, algebra_map_apply], congr, }, rw h, simp only [subalgebra.algebra_map_mem], }, - { ext f, - have h : f = 0, - { ext x', exact false.elim (n ⟨x'⟩), }, - subst h, - simp only [subalgebra.zero_mem], }, -end +end⟩ end algebra_structure diff --git a/src/topology/continuous_function/stone_weierstrass.lean b/src/topology/continuous_function/stone_weierstrass.lean index b9de96ac298f1..fe7cdff481cbe 100644 --- a/src/topology/continuous_function/stone_weierstrass.lean +++ b/src/topology/continuous_function/stone_weierstrass.lean @@ -190,11 +190,8 @@ begin and finally using compactness to produce the desired function `h` as a maximum over finitely many `x` of a minimum over finitely many `y` of the `g x y`. -/ - dsimp [set.separates_points_strongly] at sep, - - let g : X → X → L := λ x y, (sep f x y).some, - have w₁ : ∀ x y, g x y x = f x := λ x y, (sep f x y).some_spec.1, - have w₂ : ∀ x y, g x y y = f y := λ x y, (sep f x y).some_spec.2, + dsimp only [set.separates_points_strongly] at sep, + choose g hg w₁ w₂ using sep f, -- For each `x y`, we define `U x y` to be `{z | f z - ε < g x y z}`, -- and observe this is a neighbourhood of `y`. @@ -226,7 +223,7 @@ begin -- and `h x x = f x`. let h : Π x, L := λ x, ⟨(ys x).sup' (ys_nonempty x) (λ y, (g x y : C(X, ℝ))), - finset.sup'_mem _ sup_mem _ _ _ (λ y _, (g x y).2)⟩, + finset.sup'_mem _ sup_mem _ _ _ (λ y _, hg x y)⟩, have lt_h : ∀ x z, f z - ε < h x z, { intros x z, obtain ⟨y, ym, zm⟩ := set.exists_set_mem_of_union_eq_top _ _ (ys_w x) z, @@ -234,7 +231,7 @@ begin simp only [coe_fn_coe_base', subtype.coe_mk, sup'_coe, finset.sup'_apply, finset.lt_sup'_iff], exact ⟨y, ym, zm⟩ }, have h_eq : ∀ x, h x x = f x, - { intro x, simp only [coe_fn_coe_base'] at w₁, simp [coe_fn_coe_base', w₁], }, + { intro x, simp [coe_fn_coe_base', w₁], }, -- For each `x`, we define `W x` to be `{z | h x z < f z + ε}`, let W : Π x, set X := λ x, {z | h x z < f z + ε}, From f95ecd88d8ce5a6edbe90d42530cd88d48131aa4 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Fri, 28 Apr 2023 09:33:36 +0000 Subject: [PATCH 0893/1291] chore(field_theory/tower): remove unneeded imports (#18891) Co-authored-by: Scott Morrison --- docs/tutorial/representation_theory/etingof.lean | 1 + src/field_theory/tower.lean | 4 ++-- 2 files changed, 3 insertions(+), 2 deletions(-) diff --git a/docs/tutorial/representation_theory/etingof.lean b/docs/tutorial/representation_theory/etingof.lean index f56927ac0a8e2..61e788b05c04f 100644 --- a/docs/tutorial/representation_theory/etingof.lean +++ b/docs/tutorial/representation_theory/etingof.lean @@ -12,6 +12,7 @@ import algebra.category.Module.algebra import algebra.category.Module.images import algebra.category.Module.biproducts import algebra.category.Module.simple +import linear_algebra.matrix.finite_dimensional import data.mv_polynomial.basic import algebra.free_algebra import data.complex.module diff --git a/src/field_theory/tower.lean b/src/field_theory/tower.lean index 66c8e28cc56dc..3ae5c3d167246 100644 --- a/src/field_theory/tower.lean +++ b/src/field_theory/tower.lean @@ -6,8 +6,8 @@ Authors: Kenny Lau import data.nat.prime import ring_theory.algebra_tower -import linear_algebra.matrix.finite_dimensional -import linear_algebra.matrix.to_lin +import linear_algebra.finite_dimensional +import linear_algebra.free_module.finite.matrix /-! # Tower of field extensions From fa2309577c7009ea243cffdf990cd6c84f0ad497 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 28 Apr 2023 11:17:47 +0000 Subject: [PATCH 0894/1291] fix(algebra/big_operators): add missing `to_additive`s (#18878) These lemmas were written before `pi.mul_single` existed. --- src/algebra/big_operators/basic.lean | 28 ++++++++++++----------- src/algebra/big_operators/pi.lean | 34 +++++++++++++++------------- 2 files changed, 33 insertions(+), 29 deletions(-) diff --git a/src/algebra/big_operators/basic.lean b/src/algebra/big_operators/basic.lean index c84cebf6727e3..85ce482038a50 100644 --- a/src/algebra/big_operators/basic.lean +++ b/src/algebra/big_operators/basic.lean @@ -865,15 +865,16 @@ lemma prod_dite_irrel (p : Prop) [decidable p] (s : finset α) (f : p → α → (∏ x in s, if h : p then f h x else g h x) = if h : p then ∏ x in s, f h x else ∏ x in s, g h x := by { split_ifs with h; refl } -@[simp] lemma sum_pi_single' {ι M : Type*} [decidable_eq ι] [add_comm_monoid M] - (i : ι) (x : M) (s : finset ι) : - ∑ j in s, pi.single i x j = if i ∈ s then x else 0 := -sum_dite_eq' _ _ _ +@[simp, to_additive] +lemma prod_pi_mul_single' [decidable_eq α] (a : α) (x : β) (s : finset α) : + ∏ a' in s, pi.mul_single a x a' = if a ∈ s then x else 1 := +prod_dite_eq' _ _ _ -@[simp] lemma sum_pi_single {ι : Type*} {M : ι → Type*} - [decidable_eq ι] [Π i, add_comm_monoid (M i)] (i : ι) (f : Π i, M i) (s : finset ι) : - ∑ j in s, pi.single j (f j) i = if i ∈ s then f i else 0 := -sum_dite_eq _ _ _ +@[simp, to_additive] +lemma prod_pi_mul_single {β : α → Type*} + [decidable_eq α] [Π a, comm_monoid (β a)] (a : α) (f : Π a, β a) (s : finset α) : + ∏ a' in s, pi.mul_single a' (f a') a = if a ∈ s then f a else 1 := +prod_dite_eq _ _ _ @[to_additive] lemma prod_bij_ne_one {s : finset α} {t : finset γ} {f : α → β} {g : γ → β} @@ -1350,13 +1351,14 @@ begin exact λ i hi, if_neg (h i hi) } end -lemma sum_erase_lt_of_pos {γ : Type*} [decidable_eq α] [ordered_add_comm_monoid γ] - [covariant_class γ γ (+) (<)] {s : finset α} {d : α} (hd : d ∈ s) {f : α → γ} (hdf : 0 < f d) : - ∑ (m : α) in s.erase d, f m < ∑ (m : α) in s, f m := +@[to_additive] +lemma prod_erase_lt_of_one_lt {γ : Type*} [decidable_eq α] [ordered_comm_monoid γ] + [covariant_class γ γ (*) (<)] {s : finset α} {d : α} (hd : d ∈ s) {f : α → γ} (hdf : 1 < f d) : + ∏ (m : α) in s.erase d, f m < ∏ (m : α) in s, f m := begin nth_rewrite_rhs 0 ←finset.insert_erase hd, - rw finset.sum_insert (finset.not_mem_erase d s), - exact lt_add_of_pos_left _ hdf, + rw finset.prod_insert (finset.not_mem_erase d s), + exact lt_mul_of_one_lt_left' _ hdf, end /-- If a product is 1 and the function is 1 except possibly at one diff --git a/src/algebra/big_operators/pi.lean b/src/algebra/big_operators/pi.lean index 4e6021c97864a..ee2f6fc4b94c5 100644 --- a/src/algebra/big_operators/pi.lean +++ b/src/algebra/big_operators/pi.lean @@ -56,35 +56,37 @@ lemma prod_mk_prod {α β γ : Type*} [comm_monoid α] [comm_monoid β] (s : fin by haveI := classical.dec_eq γ; exact finset.induction_on s rfl (by simp [prod.ext_iff] {contextual := tt}) -section single +section mul_single variables {I : Type*} [decidable_eq I] {Z : I → Type*} -variables [Π i, add_comm_monoid (Z i)] +variables [Π i, comm_monoid (Z i)] --- As we only defined `single` into `add_monoid`, we only prove the `finset.sum` version here. -lemma finset.univ_sum_single [fintype I] (f : Π i, Z i) : - ∑ i, pi.single i (f i) = f := +@[to_additive] +lemma finset.univ_prod_mul_single [fintype I] (f : Π i, Z i) : + ∏ i, pi.mul_single i (f i) = f := by { ext a, simp } -lemma add_monoid_hom.functions_ext [finite I] (G : Type*) [add_comm_monoid G] - (g h : (Π i, Z i) →+ G) (H : ∀ i x, g (pi.single i x) = h (pi.single i x)) : g = h := +@[to_additive] +lemma monoid_hom.functions_ext [finite I] (G : Type*) [comm_monoid G] + (g h : (Π i, Z i) →* G) (H : ∀ i x, g (pi.mul_single i x) = h (pi.mul_single i x)) : g = h := begin casesI nonempty_fintype I, ext k, - rw [← finset.univ_sum_single k, g.map_sum, h.map_sum], + rw [← finset.univ_prod_mul_single k, g.map_prod, h.map_prod], simp only [H] end -/-- This is used as the ext lemma instead of `add_monoid_hom.functions_ext` for reasons explained in +/-- This is used as the ext lemma instead of `monoid_hom.functions_ext` for reasons explained in note [partially-applied ext lemmas]. -/ -@[ext] -lemma add_monoid_hom.functions_ext' [finite I] (M : Type*) [add_comm_monoid M] - (g h : (Π i, Z i) →+ M) - (H : ∀ i, g.comp (add_monoid_hom.single Z i) = h.comp (add_monoid_hom.single Z i)) : +@[ext, to_additive " +This is used as the ext lemma instead of `add_monoid_hom.functions_ext` for reasons explained in +note [partially-applied ext lemmas]."] +lemma monoid_hom.functions_ext' [finite I] (M : Type*) [comm_monoid M] + (g h : (Π i, Z i) →* M) + (H : ∀ i, g.comp (monoid_hom.single Z i) = h.comp (monoid_hom.single Z i)) : g = h := -have _ := λ i, add_monoid_hom.congr_fun (H i), -- elab without an expected type -g.functions_ext M h this +g.functions_ext M h $ λ i, monoid_hom.congr_fun (H i) -end single +end mul_single section ring_hom open pi From 3974a774a707e2e06046a14c0eaef4654584fada Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Fri, 28 Apr 2023 11:17:48 +0000 Subject: [PATCH 0895/1291] chore(category_theory/limits/preserves/finite): avoid instances with unbound universes (#18890) There were some instances with unbounded universe variables, which Lean 4 couldn't cope with (reasonably!). This PR backports some changes made in https://github.com/leanprover-community/mathlib4/pull/3615 to remove these bad instances. Co-authored-by: Scott Morrison --- .../limits/preserves/finite.lean | 31 ++++++++++++++++--- .../preadditive/injective.lean | 6 ++-- .../preadditive/projective.lean | 6 ++-- 3 files changed, 34 insertions(+), 9 deletions(-) diff --git a/src/category_theory/limits/preserves/finite.lean b/src/category_theory/limits/preserves/finite.lean index bde7a40524113..2512c091e6de6 100644 --- a/src/category_theory/limits/preserves/finite.lean +++ b/src/category_theory/limits/preserves/finite.lean @@ -55,8 +55,9 @@ noncomputable instance preserves_limits_of_shape_of_preserves_finite_limits (F : preserves_limits_of_shape J F := by apply preserves_limits_of_shape_of_equiv (fin_category.equiv_as_type J) -@[priority 100] -noncomputable instance preserves_limits.preserves_finite_limits_of_size (F : C ⥤ D) +-- This is a dangerous instance as it has unbound universe variables. +/-- If we preserve limits of some arbitrary size, then we preserve all finite limits. -/ +noncomputable def preserves_limits_of_size.preserves_finite_limits (F : C ⥤ D) [preserves_limits_of_size.{w w₂} F] : preserves_finite_limits F := ⟨λ J sJ fJ, begin @@ -64,10 +65,17 @@ noncomputable instance preserves_limits.preserves_finite_limits_of_size (F : C exact preserves_limits_of_shape_of_equiv (fin_category.equiv_as_type J) F, end⟩ +-- Added as a specialization of the dangerous instance above, for limits indexed in Type 0. +@[priority 120] +noncomputable instance preserves_limits_of_size.zero.preserves_finite_limits (F : C ⥤ D) + [preserves_limits_of_size.{0 0} F] : preserves_finite_limits F := + preserves_limits_of_size.preserves_finite_limits F + +-- An alternative specialization of the dangerous instance for small limits. @[priority 120] noncomputable instance preserves_limits.preserves_finite_limits (F : C ⥤ D) [preserves_limits F] : preserves_finite_limits F := -preserves_limits.preserves_finite_limits_of_size F +preserves_limits_of_size.preserves_finite_limits F /-- We can always derive `preserves_finite_limits C` by showing that we are preserving limits at an arbitrary universe. -/ @@ -109,8 +117,9 @@ noncomputable instance preserves_colimits_of_shape_of_preserves_finite_colimits preserves_colimits_of_shape J F := by apply preserves_colimits_of_shape_of_equiv (fin_category.equiv_as_type J) -@[priority 100] -noncomputable instance preserves_colimits.preserves_finite_colimits (F : C ⥤ D) +/-- If we preserve colimits of some arbitrary size, then we preserve all finite colimits. -/ +-- This is a dangerous instance as it has unbound universe variables. +noncomputable def preserves_colimits_of_size.preserves_finite_colimits (F : C ⥤ D) [preserves_colimits_of_size.{w w₂} F] : preserves_finite_colimits F := ⟨λ J sJ fJ, begin @@ -118,6 +127,18 @@ noncomputable instance preserves_colimits.preserves_finite_colimits (F : C ⥤ D exact preserves_colimits_of_shape_of_equiv (fin_category.equiv_as_type J) F, end⟩ +-- Added as a specialization of the dangerous instance above, for colimits indexed in Type 0. +@[priority 120] +noncomputable instance preserves_colimits_of_size.zero.preserves_finite_colimits (F : C ⥤ D) + [preserves_colimits_of_size.{0 0} F] : preserves_finite_colimits F := + preserves_colimits_of_size.preserves_finite_colimits F + +-- An alternative specialization of the dangerous instance for small colimits. +@[priority 120] +noncomputable instance preserves_colimits.preserves_finite_colimits (F : C ⥤ D) + [preserves_colimits F] : preserves_finite_colimits F := +preserves_colimits_of_size.preserves_finite_colimits F + /-- We can always derive `preserves_finite_limits C` by showing that we are preserving limits at an arbitrary universe. -/ def preserves_finite_colimits_of_preserves_finite_colimits_of_size (F : C ⥤ D) diff --git a/src/category_theory/preadditive/injective.lean b/src/category_theory/preadditive/injective.lean index b61b35c1b5b72..d791d6ecc1294 100644 --- a/src/category_theory/preadditive/injective.lean +++ b/src/category_theory/preadditive/injective.lean @@ -290,7 +290,7 @@ lemma injective_of_map_injective (adj : F ⊣ G) [full G] [faithful G] (I : D) (hI : injective (G.obj I)) : injective I := ⟨λ X Y f g, begin introI, - haveI := adj.right_adjoint_preserves_limits, + haveI : preserves_limits_of_size.{0 0} G := adj.right_adjoint_preserves_limits, rcases hI.factors (G.map f) (G.map g), use inv (adj.counit.app _) ≫ F.map w ≫ adj.counit.app _, refine faithful.map_injective G _, @@ -304,7 +304,9 @@ def map_injective_presentation (adj : F ⊣ G) [F.preserves_monomorphisms] (X : { J := G.obj I.J, injective := adj.map_injective _ I.injective, f := G.map I.f, - mono := by haveI := adj.right_adjoint_preserves_limits; apply_instance } + mono := by + haveI : preserves_limits_of_size.{0 0} G := adj.right_adjoint_preserves_limits; + apply_instance } end adjunction namespace equivalence diff --git a/src/category_theory/preadditive/projective.lean b/src/category_theory/preadditive/projective.lean index 25bd3d923a45b..6e9bbbfe779f6 100644 --- a/src/category_theory/preadditive/projective.lean +++ b/src/category_theory/preadditive/projective.lean @@ -203,7 +203,7 @@ lemma projective_of_map_projective (adj : F ⊣ G) [full F] [faithful F] (P : C) (hP : projective (F.obj P)) : projective P := ⟨λ X Y f g, begin introI, - haveI := adj.left_adjoint_preserves_colimits, + haveI : preserves_colimits_of_size.{0 0} F := adj.left_adjoint_preserves_colimits, rcases @hP.1 (F.map f) (F.map g), use adj.unit.app _ ≫ G.map w ≫ (inv $ adj.unit.app _), refine faithful.map_injective F _, @@ -217,7 +217,9 @@ def map_projective_presentation (adj : F ⊣ G) [G.preserves_epimorphisms] (X : { P := F.obj Y.P, projective := adj.map_projective _ Y.projective, f := F.map Y.f, - epi := by haveI := adj.left_adjoint_preserves_colimits; apply_instance } + epi := by + haveI : preserves_colimits_of_size.{0 0} F := adj.left_adjoint_preserves_colimits; + apply_instance } end adjunction namespace equivalence From ef093414afad9796939469de466cc3c206e18223 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 28 Apr 2023 14:14:13 +0000 Subject: [PATCH 0896/1291] chore(measure_theory/measure/haar_of_basis): axis-aligned parallepipeds are cuboids (#18829) This was particularly annoying to prove; it feels like I might be missing some API somewhere. This also includes a result that @YaelDillies proved for me that shows a parallepiped can be formed by summing all the elements on its primary edges. --- src/measure_theory/measure/haar_lebesgue.lean | 9 +++ src/measure_theory/measure/haar_of_basis.lean | 72 ++++++++++++++++++- 2 files changed, 80 insertions(+), 1 deletion(-) diff --git a/src/measure_theory/measure/haar_lebesgue.lean b/src/measure_theory/measure/haar_lebesgue.lean index ac4e7f454a807..965c33384fa36 100644 --- a/src/measure_theory/measure/haar_lebesgue.lean +++ b/src/measure_theory/measure/haar_lebesgue.lean @@ -60,6 +60,15 @@ def topological_space.positive_compacts.pi_Icc01 (ι : Type*) [fintype ι] : interior_nonempty' := by simp only [interior_pi_set, set.to_finite, interior_Icc, univ_pi_nonempty_iff, nonempty_Ioo, implies_true_iff, zero_lt_one] } +/-- The parallelepiped formed from the standard basis for `ι → ℝ` is `[0,1]^ι` -/ +lemma basis.parallelepiped_basis_fun (ι : Type*) [fintype ι] : + (pi.basis_fun ℝ ι).parallelepiped = topological_space.positive_compacts.pi_Icc01 ι := +set_like.coe_injective $ begin + refine eq.trans _ ((uIcc_of_le _).trans (set.pi_univ_Icc _ _).symm), + { convert (parallelepiped_single 1) }, + { exact zero_le_one }, +end + namespace measure_theory open measure topological_space.positive_compacts finite_dimensional diff --git a/src/measure_theory/measure/haar_of_basis.lean b/src/measure_theory/measure/haar_of_basis.lean index cbb1fe7460ded..ed046f5070061 100644 --- a/src/measure_theory/measure/haar_of_basis.lean +++ b/src/measure_theory/measure/haar_of_basis.lean @@ -27,7 +27,7 @@ of the basis). -/ open set topological_space measure_theory measure_theory.measure finite_dimensional -open_locale big_operators +open_locale big_operators pointwise noncomputable theory @@ -106,6 +106,76 @@ begin neg_zero, finset.univ_unique] }, end +lemma parallelepiped_eq_sum_segment (v : ι → E) : parallelepiped v = ∑ i, segment ℝ 0 (v i) := +begin + ext, + simp only [mem_parallelepiped_iff, set.mem_finset_sum, finset.mem_univ, forall_true_left, + segment_eq_image, smul_zero, zero_add, ←set.pi_univ_Icc, set.mem_univ_pi], + split, + { rintro ⟨t, ht, rfl⟩, + exact ⟨t • v, λ i, ⟨t i, ht _, by simp⟩, rfl⟩ }, + rintro ⟨g, hg, rfl⟩, + change ∀ i, _ at hg, + choose t ht hg using hg, + refine ⟨t, ht, _⟩, + simp_rw hg, +end + +lemma convex_parallelepiped (v : ι → E) : convex ℝ (parallelepiped v) := +begin + rw parallelepiped_eq_sum_segment, + -- TODO: add `convex.sum` to match `convex.add` + let : add_submonoid (set E) := + { carrier := { s | convex ℝ s}, zero_mem' := convex_singleton _, add_mem' := λ x y, convex.add }, + exact this.sum_mem (λ i hi, convex_segment _ _), +end + +/-- A `parallelepiped` is the convex hull of its vertices -/ +lemma parallelepiped_eq_convex_hull (v : ι → E) : + parallelepiped v = convex_hull ℝ (∑ i, {(0 : E), v i}) := +begin + -- TODO: add `convex_hull_sum` to match `convex_hull_add` + let : set E →+ set E := + { to_fun := convex_hull ℝ, + map_zero' := convex_hull_singleton _, + map_add' := convex_hull_add }, + simp_rw [parallelepiped_eq_sum_segment, ←convex_hull_pair], + exact (this.map_sum _ _).symm, +end + +/-- The axis aligned parallelepiped over `ι → ℝ` is a cuboid. -/ +lemma parallelepiped_single [decidable_eq ι] (a : ι → ℝ) : + parallelepiped (λ i, pi.single i (a i)) = set.uIcc 0 a := +begin + ext, + simp_rw [set.uIcc, mem_parallelepiped_iff, set.mem_Icc, pi.le_def, ←forall_and_distrib, + pi.inf_apply, pi.sup_apply, ←pi.single_smul', pi.one_apply, pi.zero_apply, ←pi.smul_apply', + finset.univ_sum_single (_ : ι → ℝ)], + split, + { rintros ⟨t, ht, rfl⟩ i, + specialize ht i, + simp_rw [smul_eq_mul, pi.mul_apply], + cases le_total (a i) 0 with hai hai, + { rw [sup_eq_left.mpr hai, inf_eq_right.mpr hai], + exact ⟨le_mul_of_le_one_left hai ht.2, mul_nonpos_of_nonneg_of_nonpos ht.1 hai⟩ }, + { rw [sup_eq_right.mpr hai, inf_eq_left.mpr hai], + exact ⟨mul_nonneg ht.1 hai, mul_le_of_le_one_left hai ht.2⟩ } }, + { intro h, + refine ⟨λ i, x i / a i, λ i, _, funext $ λ i, _⟩, + { specialize h i, + cases le_total (a i) 0 with hai hai, + { rw [sup_eq_left.mpr hai, inf_eq_right.mpr hai] at h, + exact ⟨div_nonneg_of_nonpos h.2 hai, div_le_one_of_ge h.1 hai⟩ }, + { rw [sup_eq_right.mpr hai, inf_eq_left.mpr hai] at h, + exact ⟨div_nonneg h.1 hai, div_le_one_of_le h.2 hai⟩ } }, + { specialize h i, + simp only [smul_eq_mul, pi.mul_apply], + cases eq_or_ne (a i) 0 with hai hai, + { rw [hai, inf_idem, sup_idem, ←le_antisymm_iff] at h, + rw [hai, ← h, zero_div, zero_mul] }, + { rw div_mul_cancel _ hai } } }, +end + end add_comm_group section normed_space From 9c5398f2ded9f4ff733d3c7e2c90457b943fc4fc Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Fri, 28 Apr 2023 14:14:14 +0000 Subject: [PATCH 0897/1291] feat(measure_theory/constructions/borel_space): add measurable_c{Inf/supr/infi} (#18882) `measurable_cSup` was already there. --- .../constructions/borel_space.lean | 19 +++++++++++++++++++ 1 file changed, 19 insertions(+) diff --git a/src/measure_theory/constructions/borel_space.lean b/src/measure_theory/constructions/borel_space.lean index 426ffacdb82d2..e8067d5ceca54 100644 --- a/src/measure_theory/constructions/borel_space.lean +++ b/src/measure_theory/constructions/borel_space.lean @@ -1279,6 +1279,25 @@ begin exact measurable_set.bInter hs (λ i hi, measurable_set_le (hf i) measurable_const) } end +lemma measurable_cInf {ι} {f : ι → δ → α} {s : set ι} (hs : s.countable) + (hf : ∀ i, measurable (f i)) (bdd : ∀ x, bdd_below ((λ i, f i x) '' s)) : + measurable (λ x, Inf ((λ i, f i x) '' s)) := +@measurable_cSup αᵒᵈ _ _ _ _ _ _ _ _ _ _ _ hs hf bdd + +lemma measurable_csupr {ι : Type*} [countable ι] {f : ι → δ → α} + (hf : ∀ i, measurable (f i)) (bdd : ∀ x, bdd_above (range (λ i, f i x))) : + measurable (λ x, ⨆ i, f i x) := +begin + change measurable (λ x, Sup (range (λ i : ι, f i x))), + simp_rw ← image_univ at bdd ⊢, + refine measurable_cSup countable_univ hf bdd, +end + +lemma measurable_cinfi {ι : Type*} [countable ι] {f : ι → δ → α} + (hf : ∀ i, measurable (f i)) (bdd : ∀ x, bdd_below (range (λ i, f i x))) : + measurable (λ x, ⨅ i, f i x) := +@measurable_csupr αᵒᵈ _ _ _ _ _ _ _ _ _ _ _ hf bdd + end conditionally_complete_linear_order /-- Convert a `homeomorph` to a `measurable_equiv`. -/ From f231b9d8ce4970789c592af9508e06a0884f72d1 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Fri, 28 Apr 2023 14:14:16 +0000 Subject: [PATCH 0898/1291] chore(measure_theory/integral/lebesgue): generalize lintegral_sub and lintegral_sub_le to ae_measurable (#18883) --- src/measure_theory/integral/lebesgue.lean | 17 +++++++++++++---- 1 file changed, 13 insertions(+), 4 deletions(-) diff --git a/src/measure_theory/integral/lebesgue.lean b/src/measure_theory/integral/lebesgue.lean index 6b78bb4a2bd42..60c905b776f42 100644 --- a/src/measure_theory/integral/lebesgue.lean +++ b/src/measure_theory/integral/lebesgue.lean @@ -1780,26 +1780,35 @@ calc ... = ⨆n, (∫⁻ a, f n a ∂μ) : by simp only [lintegral_congr_ae (g_eq_f.mono $ λ a ha, ha _)] -lemma lintegral_sub {f g : α → ℝ≥0∞} (hg : measurable g) +lemma lintegral_sub' {f g : α → ℝ≥0∞} (hg : ae_measurable g μ) (hg_fin : ∫⁻ a, g a ∂μ ≠ ∞) (h_le : g ≤ᵐ[μ] f) : ∫⁻ a, f a - g a ∂μ = ∫⁻ a, f a ∂μ - ∫⁻ a, g a ∂μ := begin refine ennreal.eq_sub_of_add_eq hg_fin _, - rw [← lintegral_add_right _ hg], + rw [← lintegral_add_right' _ hg], exact lintegral_congr_ae (h_le.mono $ λ x hx, tsub_add_cancel_of_le hx) end -lemma lintegral_sub_le (f g : α → ℝ≥0∞) (hf : measurable f) : +lemma lintegral_sub {f g : α → ℝ≥0∞} (hg : measurable g) + (hg_fin : ∫⁻ a, g a ∂μ ≠ ∞) (h_le : g ≤ᵐ[μ] f) : + ∫⁻ a, f a - g a ∂μ = ∫⁻ a, f a ∂μ - ∫⁻ a, g a ∂μ := +lintegral_sub' hg.ae_measurable hg_fin h_le + +lemma lintegral_sub_le' (f g : α → ℝ≥0∞) (hf : ae_measurable f μ) : ∫⁻ x, g x ∂μ - ∫⁻ x, f x ∂μ ≤ ∫⁻ x, g x - f x ∂μ := begin rw tsub_le_iff_right, by_cases hfi : ∫⁻ x, f x ∂μ = ∞, { rw [hfi, add_top], exact le_top }, - { rw [← lintegral_add_right _ hf], + { rw [← lintegral_add_right' _ hf], exact lintegral_mono (λ x, le_tsub_add) } end +lemma lintegral_sub_le (f g : α → ℝ≥0∞) (hf : measurable f) : + ∫⁻ x, g x ∂μ - ∫⁻ x, f x ∂μ ≤ ∫⁻ x, g x - f x ∂μ := +lintegral_sub_le' f g hf.ae_measurable + lemma lintegral_strict_mono_of_ae_le_of_frequently_ae_lt {f g : α → ℝ≥0∞} (hg : ae_measurable g μ) (hfi : ∫⁻ x, f x ∂μ ≠ ∞) (h_le : f ≤ᵐ[μ] g) (h : ∃ᵐ x ∂μ, f x ≠ g x) : ∫⁻ x, f x ∂μ < ∫⁻ x, g x ∂μ := From c7bce2818663f456335892ddbdd1809f111a5b72 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 28 Apr 2023 14:14:17 +0000 Subject: [PATCH 0899/1291] feat(field_theory/tower): generalize `rank_mul_rank` to rings (#18885) This also renames `rank_mul_rank` to `lift_rank_mul_lift_rank` --- src/data/complex/module.lean | 2 +- src/field_theory/tower.lean | 79 ++++++++++++++++++++++++------------ 2 files changed, 54 insertions(+), 27 deletions(-) diff --git a/src/data/complex/module.lean b/src/data/complex/module.lean index b17fd2e2a7087..b6a04214d576e 100644 --- a/src/data/complex/module.lean +++ b/src/data/complex/module.lean @@ -198,7 +198,7 @@ finite_dimensional.trans ℝ ℂ E lemma rank_real_of_complex (E : Type*) [add_comm_group E] [module ℂ E] : module.rank ℝ E = 2 * module.rank ℂ E := cardinal.lift_inj.1 $ - by { rw [← rank_mul_rank' ℝ ℂ E, complex.rank_real_complex], simp [bit0] } + by { rw [← lift_rank_mul_lift_rank ℝ ℂ E, complex.rank_real_complex], simp [bit0] } lemma finrank_real_of_complex (E : Type*) [add_comm_group E] [module ℂ E] : finite_dimensional.finrank ℝ E = 2 * finite_dimensional.finrank ℂ E := diff --git a/src/field_theory/tower.lean b/src/field_theory/tower.lean index 3ae5c3d167246..f35f35b430c1d 100644 --- a/src/field_theory/tower.lean +++ b/src/field_theory/tower.lean @@ -16,8 +16,8 @@ In this file we prove the tower law for arbitrary extensions and finite extensio Suppose `L` is a field extension of `K` and `K` is a field extension of `F`. Then `[L:F] = [L:K] [K:F]` where `[E₁:E₂]` means the `E₂`-dimension of `E₁`. -In fact we generalize it to vector spaces, where `L` is not necessarily a field, -but just a vector space over `K`. +In fact we generalize it to rings and modules, where `L` is not necessarily a field, +but just a free module over `K`. ## Implementation notes @@ -34,31 +34,60 @@ tower law universes u v w u₁ v₁ w₁ open_locale classical big_operators -section field - +open finite_dimensional open cardinal variables (F : Type u) (K : Type v) (A : Type w) -variables [field F] [division_ring K] [add_comm_group A] + +section ring + +variables [comm_ring F] [ring K] [add_comm_group A] variables [algebra F K] [module K A] [module F A] [is_scalar_tower F K A] +variables [strong_rank_condition F] [strong_rank_condition K] [module.free F K] [module.free K A] -/-- Tower law: if `A` is a `K`-vector space and `K` is a field extension of `F` then -`dim_F(A) = dim_F(K) * dim_K(A)`. -/ -theorem rank_mul_rank' : +/-- Tower law: if `A` is a `K`-module and `K` is an extension of `F` then +$\operatorname{rank}_F(A) = \operatorname{rank}_F(K) * \operatorname{rank}_K(A)$. -/ +theorem lift_rank_mul_lift_rank : (cardinal.lift.{w} (module.rank F K) * cardinal.lift.{v} (module.rank K A)) = cardinal.lift.{v} (module.rank F A) := -let b := basis.of_vector_space F K, c := basis.of_vector_space K A in -by rw [← (module.rank F K).lift_id, ← b.mk_eq_rank, - ← (module.rank K A).lift_id, ← c.mk_eq_rank, - ← lift_umax.{w v}, ← (b.smul c).mk_eq_rank, mk_prod, lift_mul, - lift_lift, lift_lift, lift_lift, lift_lift, lift_umax] +begin + obtain ⟨_, b⟩ := module.free.exists_basis F K, + obtain ⟨_, c⟩ := module.free.exists_basis K A, + rw [← (module.rank F K).lift_id, ← b.mk_eq_rank, + ← (module.rank K A).lift_id, ← c.mk_eq_rank, + ← lift_umax.{w v}, ← (b.smul c).mk_eq_rank, mk_prod, lift_mul, + lift_lift, lift_lift, lift_lift, lift_lift, lift_umax] +end -/-- Tower law: if `A` is a `K`-vector space and `K` is a field extension of `F` then -`dim_F(A) = dim_F(K) * dim_K(A)`. -/ -theorem rank_mul_rank (F : Type u) (K A : Type v) [field F] [field K] [add_comm_group A] - [algebra F K] [module K A] [module F A] [is_scalar_tower F K A] : +/-- Tower law: if `A` is a `K`-module and `K` is an extension of `F` then +$\operatorname{rank}_F(A) = \operatorname{rank}_F(K) * \operatorname{rank}_K(A)$. + +This is a simpler version of `lift_rank_mul_lift_rank` with `K` and `A` in the same universe. -/ +theorem rank_mul_rank (F : Type u) (K A : Type v) + [comm_ring F] [ring K] [add_comm_group A] + [algebra F K] [module K A] [module F A] [is_scalar_tower F K A] + [strong_rank_condition F] [strong_rank_condition K] [module.free F K] [module.free K A] : module.rank F K * module.rank K A = module.rank F A := -by convert rank_mul_rank' F K A; rw lift_id +by convert lift_rank_mul_lift_rank F K A; rw lift_id + +/-- Tower law: if `A` is a `K`-module and `K` is an extension of `F` then +$\operatorname{rank}_F(A) = \operatorname{rank}_F(K) * \operatorname{rank}_K(A)$. -/ +theorem finite_dimensional.finrank_mul_finrank' + [nontrivial K] [module.finite F K] [module.finite K A] : + finrank F K * finrank K A = finrank F A := +begin + letI := nontrivial_of_invariant_basis_number F, + let b := module.free.choose_basis F K, + let c := module.free.choose_basis K A, + rw [finrank_eq_card_basis b, finrank_eq_card_basis c, finrank_eq_card_basis (b.smul c), + fintype.card_prod], +end + +end ring + +section field +variables [field F] [division_ring K] [add_comm_group A] +variables [algebra F K] [module K A] [module F A] [is_scalar_tower F K A] namespace finite_dimensional @@ -85,17 +114,15 @@ let ⟨⟨b, hb⟩⟩ := hf in ⟨⟨b, submodule.restrict_scalars_injective F _ by { rw [submodule.restrict_scalars_top, eq_top_iff, ← hb, submodule.span_le], exact submodule.subset_span }⟩⟩ -/-- Tower law: if `A` is a `K`-algebra and `K` is a field extension of `F` then -`dim_F(A) = dim_F(K) * dim_K(A)`. -/ -theorem finrank_mul_finrank [finite_dimensional F K] : - finrank F K * finrank K A = finrank F A := +/-- Tower law: if `A` is a `K`-vector space and `K` is a field extension of `F` then +`dim_F(A) = dim_F(K) * dim_K(A)`. + +This is `finite_dimensional.finrank_mul_finrank'` with one fewer finiteness assumption. -/ +theorem finrank_mul_finrank [finite_dimensional F K] : finrank F K * finrank K A = finrank F A := begin by_cases hA : finite_dimensional K A, { resetI, - let b := basis.of_vector_space F K, - let c := basis.of_vector_space K A, - rw [finrank_eq_card_basis b, finrank_eq_card_basis c, - finrank_eq_card_basis (b.smul c), fintype.card_prod] }, + rw finrank_mul_finrank' }, { rw [finrank_of_infinite_dimensional hA, mul_zero, finrank_of_infinite_dimensional], exact mt (@right F K A _ _ _ _ _ _ _) hA } end From c813ed7de0f5115f956239124e9b30f3a621966f Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 28 Apr 2023 14:14:18 +0000 Subject: [PATCH 0900/1291] chore(data/finset/lattice): Remove `finset.sup_finset_image` (#18893) in favor of the identical `finset.sup_image` (up to argument order) in the same file. --- src/data/finset/lattice.lean | 8 -------- src/order/compactly_generated.lean | 2 +- src/ring_theory/finiteness.lean | 2 +- src/ring_theory/mv_polynomial/symmetric.lean | 2 +- 4 files changed, 3 insertions(+), 11 deletions(-) diff --git a/src/data/finset/lattice.lean b/src/data/finset/lattice.lean index c44107febc9ca..828b64da31613 100644 --- a/src/data/finset/lattice.lean +++ b/src/data/finset/lattice.lean @@ -1479,14 +1479,6 @@ lemma supr_finset_image {f : γ → α} {g : α → β} {s : finset γ} : (⨆ x ∈ s.image f, g x) = (⨆ y ∈ s, g (f y)) := by rw [← supr_coe, coe_image, supr_image, supr_coe] -lemma sup_finset_image {β γ : Type*} [semilattice_sup β] [order_bot β] - (f : γ → α) (g : α → β) (s : finset γ) : - (s.image f).sup g = s.sup (g ∘ f) := -begin - classical, - induction s using finset.induction_on with a s' ha ih; simp * -end - lemma infi_finset_image {f : γ → α} {g : α → β} {s : finset γ} : (⨅ x ∈ s.image f, g x) = (⨅ y ∈ s, g (f y)) := by rw [← infi_coe, coe_image, infi_image, infi_coe] diff --git a/src/order/compactly_generated.lean b/src/order/compactly_generated.lean index bd288f7693805..f08caf837ce5a 100644 --- a/src/order/compactly_generated.lean +++ b/src/order/compactly_generated.lean @@ -182,7 +182,7 @@ begin classical, rw is_compact_element_iff_le_of_directed_Sup_le, intros d hemp hdir hsup, - change f with id ∘ f, rw ←finset.sup_finset_image, + change f with id ∘ f, rw ←finset.sup_image, apply finset.sup_le_of_le_directed d hemp hdir, rintros x hx, obtain ⟨p, ⟨hps, rfl⟩⟩ := finset.mem_image.mp hx, diff --git a/src/ring_theory/finiteness.lean b/src/ring_theory/finiteness.lean index a337c43640ef1..770b1bb947c6e 100644 --- a/src/ring_theory/finiteness.lean +++ b/src/ring_theory/finiteness.lean @@ -345,7 +345,7 @@ begin { suffices : u.sup id ≤ s, from le_antisymm husup this, rw [sSup, finset.sup_id_eq_Sup], exact Sup_le_Sup huspan, }, obtain ⟨t, ⟨hts, rfl⟩⟩ := finset.subset_image_iff.mp huspan, - rw [finset.sup_finset_image, function.comp.left_id, finset.sup_eq_supr, supr_rw, + rw [finset.sup_image, function.comp.left_id, finset.sup_eq_supr, supr_rw, ←span_eq_supr_of_singleton_spans, eq_comm] at ssup, exact ⟨t, ssup⟩, }, end diff --git a/src/ring_theory/mv_polynomial/symmetric.lean b/src/ring_theory/mv_polynomial/symmetric.lean index 4922a8cfafbc2..d0ab30e3b7176 100644 --- a/src/ring_theory/mv_polynomial/symmetric.lean +++ b/src/ring_theory/mv_polynomial/symmetric.lean @@ -215,7 +215,7 @@ begin classical, have : (finsupp.to_multiset ∘ λ (t : finset σ), ∑ (i : σ) in t, finsupp.single i 1) = finset.val, { funext, simp [finsupp.to_multiset_sum_single] }, - rw [degrees, support_esymm, sup_finset_image, this, ←comp_sup_eq_sup_comp], + rw [degrees, support_esymm, sup_image, this, ←comp_sup_eq_sup_comp], { obtain ⟨k, rfl⟩ := nat.exists_eq_succ_of_ne_zero hpos.ne', simpa using powerset_len_sup _ _ (nat.lt_of_succ_le hn) }, { intros, From 84771a9f5f0bd5e5d6218811556508ddf476dcbd Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 28 Apr 2023 17:05:16 +0000 Subject: [PATCH 0901/1291] =?UTF-8?q?feat(algebra/group/basic):=20`a=20*?= =?UTF-8?q?=20b=20=E2=89=A0=20b=20=E2=86=94=20a=20=E2=89=A0=201`=20(#18635?= =?UTF-8?q?)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit A few convenient corollaries of existing lemmas and `a / 2 ≤ a ↔ 0 ≤ a`. --- src/algebra/group/basic.lean | 6 ++++++ src/algebra/order/field/basic.lean | 17 +++++++---------- 2 files changed, 13 insertions(+), 10 deletions(-) diff --git a/src/algebra/group/basic.lean b/src/algebra/group/basic.lean index 04c71bc61da11..bff1c7e7278ad 100644 --- a/src/algebra/group/basic.lean +++ b/src/algebra/group/basic.lean @@ -159,6 +159,9 @@ calc a * b = a ↔ a * b = a * 1 : by rw mul_one @[simp, to_additive] lemma self_eq_mul_right : a = a * b ↔ b = 1 := eq_comm.trans mul_right_eq_self +@[to_additive] lemma mul_right_ne_self : a * b ≠ a ↔ b ≠ 1 := mul_right_eq_self.not +@[to_additive] lemma self_ne_mul_right : a ≠ a * b ↔ b ≠ 1 := self_eq_mul_right.not + end left_cancel_monoid section right_cancel_monoid @@ -172,6 +175,9 @@ calc a * b = b ↔ a * b = 1 * b : by rw one_mul @[simp, to_additive] lemma self_eq_mul_left : b = a * b ↔ a = 1 := eq_comm.trans mul_left_eq_self +@[to_additive] lemma mul_left_ne_self : a * b ≠ b ↔ a ≠ 1 := mul_left_eq_self.not +@[to_additive] lemma self_ne_mul_left : b ≠ a * b ↔ a ≠ 1 := self_eq_mul_left.not + end right_cancel_monoid section has_involutive_inv diff --git a/src/algebra/order/field/basic.lean b/src/algebra/order/field/basic.lean index 3477d93077b99..049eff237f9b8 100644 --- a/src/algebra/order/field/basic.lean +++ b/src/algebra/order/field/basic.lean @@ -407,18 +407,15 @@ lemma half_pos (h : 0 < a) : 0 < a / 2 := div_pos h zero_lt_two lemma one_half_pos : (0:α) < 1 / 2 := half_pos zero_lt_one -lemma div_two_lt_of_pos (h : 0 < a) : a / 2 < a := -by { rw [div_lt_iff (zero_lt_two' α)], exact lt_mul_of_one_lt_right h one_lt_two } +@[simp] lemma half_le_self_iff : a / 2 ≤ a ↔ 0 ≤ a := +by rw [div_le_iff (zero_lt_two' α), mul_two, le_add_iff_nonneg_left] -lemma half_lt_self : 0 < a → a / 2 < a := div_two_lt_of_pos +@[simp] lemma half_lt_self_iff : a / 2 < a ↔ 0 < a := +by rw [div_lt_iff (zero_lt_two' α), mul_two, lt_add_iff_pos_left] -lemma half_le_self (ha_nonneg : 0 ≤ a) : a / 2 ≤ a := -begin - by_cases h0 : a = 0, - { simp [h0], }, - { rw ← ne.def at h0, - exact (half_lt_self (lt_of_le_of_ne ha_nonneg h0.symm)).le, }, -end +alias half_le_self_iff ↔ _ half_le_self +alias half_lt_self_iff ↔ _ half_lt_self +alias half_lt_self ← div_two_lt_of_pos lemma one_half_lt_one : (1 / 2 : α) < 1 := half_lt_self zero_lt_one From eee67e6e5fb178ca03a5880d7b8c509e6f0b868b Mon Sep 17 00:00:00 2001 From: Alex J Best Date: Fri, 28 Apr 2023 17:05:17 +0000 Subject: [PATCH 0902/1291] chore(.github/workflows): update cancel workflow action (#18761) To stop deprecation warnings on every PR. The test commit shows that this works as intended --- .github/workflows/bors.yml | 2 +- .github/workflows/build.yml | 2 +- .github/workflows/build.yml.in | 2 +- .github/workflows/build_fork.yml | 2 +- .github/workflows/dependent-issues.yml | 2 +- 5 files changed, 5 insertions(+), 5 deletions(-) diff --git a/.github/workflows/bors.yml b/.github/workflows/bors.yml index 0b8033c11e328..e0f6e696b7e8e 100644 --- a/.github/workflows/bors.yml +++ b/.github/workflows/bors.yml @@ -22,7 +22,7 @@ jobs: runs-on: ubuntu-latest # timeout-minutes: 3 steps: - - uses: styfle/cancel-workflow-action@0.9.0 + - uses: styfle/cancel-workflow-action@0.11.0 with: all_but_latest: true access_token: ${{ github.token }} diff --git a/.github/workflows/build.yml b/.github/workflows/build.yml index b61b8eea978f3..b7fc0f16f8a37 100644 --- a/.github/workflows/build.yml +++ b/.github/workflows/build.yml @@ -30,7 +30,7 @@ jobs: runs-on: ubuntu-latest # timeout-minutes: 3 steps: - - uses: styfle/cancel-workflow-action@0.9.0 + - uses: styfle/cancel-workflow-action@0.11.0 with: all_but_latest: true access_token: ${{ github.token }} diff --git a/.github/workflows/build.yml.in b/.github/workflows/build.yml.in index a3f6edde6168c..324bbeeb942c6 100644 --- a/.github/workflows/build.yml.in +++ b/.github/workflows/build.yml.in @@ -8,7 +8,7 @@ jobs: runs-on: ubuntu-latest # timeout-minutes: 3 steps: - - uses: styfle/cancel-workflow-action@0.9.0 + - uses: styfle/cancel-workflow-action@0.11.0 with: all_but_latest: true access_token: ${{ github.token }} diff --git a/.github/workflows/build_fork.yml b/.github/workflows/build_fork.yml index 91ad14c848268..2bd1e4e2da4ad 100644 --- a/.github/workflows/build_fork.yml +++ b/.github/workflows/build_fork.yml @@ -28,7 +28,7 @@ jobs: runs-on: ubuntu-latest # timeout-minutes: 3 steps: - - uses: styfle/cancel-workflow-action@0.9.0 + - uses: styfle/cancel-workflow-action@0.11.0 with: all_but_latest: true access_token: ${{ github.token }} diff --git a/.github/workflows/dependent-issues.yml b/.github/workflows/dependent-issues.yml index 0efb065dc44c4..682c5311f52b0 100644 --- a/.github/workflows/dependent-issues.yml +++ b/.github/workflows/dependent-issues.yml @@ -10,7 +10,7 @@ jobs: runs-on: ubuntu-latest # timeout-minutes: 3 steps: - - uses: styfle/cancel-workflow-action@0.9.0 + - uses: styfle/cancel-workflow-action@0.11.0 with: all_but_latest: true access_token: ${{ github.token }} From a4f99eae998680d3a2c240da4a2b16354c85ee49 Mon Sep 17 00:00:00 2001 From: Alex J Best Date: Fri, 28 Apr 2023 17:05:18 +0000 Subject: [PATCH 0903/1291] chore(.github/workflows): replace set-output commands (#18762) Following deprecation, c.f. https://github.blog/changelog/2022-10-11-github-actions-deprecating-save-state-and-set-output-commands/ The check that this change worked ok is that the action produces a correctly named artifact --- .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 e0f6e696b7e8e..b35eb1d373d3d 100644 --- a/.github/workflows/bors.yml +++ b/.github/workflows/bors.yml @@ -90,7 +90,7 @@ jobs: run: | set -o pipefail leanpkg configure - echo "::set-output name=started::true" + echo "started=true" >> $GITHUB_OUTPUT lean --json -T100000 --make src | python3 scripts/detect_errors.py lean --json -T400000 --make src | python3 scripts/detect_errors.py @@ -111,7 +111,7 @@ jobs: touch workspace.tar tar -cf workspace.tar --exclude=*.tar* . git_hash="$(git log -1 --pretty=format:%h)" - echo "::set-output name=artifact_name::precompiled-mathlib-$short_lean_version-$git_hash" + echo "artifact_name=precompiled-mathlib-$short_lean_version-$git_hash" >> $GITHUB_OUTPUT - name: upload precompiled mathlib zip file if: always() && steps.build.outputs.started == 'true' diff --git a/.github/workflows/build.yml b/.github/workflows/build.yml index b7fc0f16f8a37..f39161df1a753 100644 --- a/.github/workflows/build.yml +++ b/.github/workflows/build.yml @@ -98,7 +98,7 @@ jobs: run: | set -o pipefail leanpkg configure - echo "::set-output name=started::true" + echo "started=true" >> $GITHUB_OUTPUT lean --json -T100000 --make src | python3 scripts/detect_errors.py lean --json -T400000 --make src | python3 scripts/detect_errors.py @@ -119,7 +119,7 @@ jobs: touch workspace.tar tar -cf workspace.tar --exclude=*.tar* . git_hash="$(git log -1 --pretty=format:%h)" - echo "::set-output name=artifact_name::precompiled-mathlib-$short_lean_version-$git_hash" + echo "artifact_name=precompiled-mathlib-$short_lean_version-$git_hash" >> $GITHUB_OUTPUT - name: upload precompiled mathlib zip file if: always() && steps.build.outputs.started == 'true' diff --git a/.github/workflows/build.yml.in b/.github/workflows/build.yml.in index 324bbeeb942c6..d524f35872bb5 100644 --- a/.github/workflows/build.yml.in +++ b/.github/workflows/build.yml.in @@ -76,7 +76,7 @@ jobs: run: | set -o pipefail leanpkg configure - echo "::set-output name=started::true" + echo "started=true" >> $GITHUB_OUTPUT lean --json -T100000 --make src | python3 scripts/detect_errors.py lean --json -T400000 --make src | python3 scripts/detect_errors.py @@ -97,7 +97,7 @@ jobs: touch workspace.tar tar -cf workspace.tar --exclude=*.tar* . git_hash="$(git log -1 --pretty=format:%h)" - echo "::set-output name=artifact_name::precompiled-mathlib-$short_lean_version-$git_hash" + echo "artifact_name=precompiled-mathlib-$short_lean_version-$git_hash" >> $GITHUB_OUTPUT - name: upload precompiled mathlib zip file if: always() && steps.build.outputs.started == 'true' diff --git a/.github/workflows/build_fork.yml b/.github/workflows/build_fork.yml index 2bd1e4e2da4ad..782a118a88ba0 100644 --- a/.github/workflows/build_fork.yml +++ b/.github/workflows/build_fork.yml @@ -96,7 +96,7 @@ jobs: run: | set -o pipefail leanpkg configure - echo "::set-output name=started::true" + echo "started=true" >> $GITHUB_OUTPUT lean --json -T100000 --make src | python3 scripts/detect_errors.py lean --json -T400000 --make src | python3 scripts/detect_errors.py @@ -117,7 +117,7 @@ jobs: touch workspace.tar tar -cf workspace.tar --exclude=*.tar* . git_hash="$(git log -1 --pretty=format:%h)" - echo "::set-output name=artifact_name::precompiled-mathlib-$short_lean_version-$git_hash" + echo "artifact_name=precompiled-mathlib-$short_lean_version-$git_hash" >> $GITHUB_OUTPUT - name: upload precompiled mathlib zip file if: always() && steps.build.outputs.started == 'true' From f2b757fc5c341d88741b9c4630b1e8ba973c5726 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Sat, 29 Apr 2023 05:34:54 +0000 Subject: [PATCH 0904/1291] chore(*): add mathlib4 synchronization comments (#18889) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebraic_topology.nerve` * `algebraic_topology.simplicial_set` * `analysis.convex.stone_separation` * `category_theory.action` * `category_theory.limits.filtered_colimit_commutes_finite_limit` * `category_theory.noetherian` * `category_theory.preadditive.projective` * `category_theory.simple` * `category_theory.sites.adjunction` * `linear_algebra.free_module.finite.matrix` * `linear_algebra.matrix.basis` * `linear_algebra.matrix.block` * `linear_algebra.matrix.diagonal` * `linear_algebra.matrix.finite_dimensional` * `linear_algebra.matrix.invariant_basis_number` * `linear_algebra.matrix.to_lin` * `linear_algebra.matrix.transvection` * `linear_algebra.matrix.zpow` * `topology.algebra.module.multilinear` * `topology.algebra.module.strong_topology` * `topology.algebra.uniform_convergence` * `topology.category.CompHaus.basic` * `topology.category.CompHaus.projective` * `topology.category.Top.limits.basic` * `topology.category.Top.opens` --- src/algebraic_topology/nerve.lean | 3 +++ src/algebraic_topology/simplicial_set.lean | 3 +++ src/analysis/convex/stone_separation.lean | 3 +++ src/category_theory/action.lean | 3 +++ .../limits/filtered_colimit_commutes_finite_limit.lean | 3 +++ src/category_theory/noetherian.lean | 3 +++ src/category_theory/preadditive/projective.lean | 3 +++ src/category_theory/simple.lean | 3 +++ src/category_theory/sites/adjunction.lean | 3 +++ src/linear_algebra/free_module/finite/matrix.lean | 3 +++ src/linear_algebra/matrix/basis.lean | 3 +++ src/linear_algebra/matrix/block.lean | 3 +++ src/linear_algebra/matrix/diagonal.lean | 3 +++ src/linear_algebra/matrix/finite_dimensional.lean | 3 +++ src/linear_algebra/matrix/invariant_basis_number.lean | 3 +++ src/linear_algebra/matrix/to_lin.lean | 3 +++ src/linear_algebra/matrix/transvection.lean | 3 +++ src/linear_algebra/matrix/zpow.lean | 3 +++ src/topology/algebra/module/multilinear.lean | 3 +++ src/topology/algebra/module/strong_topology.lean | 3 +++ src/topology/algebra/uniform_convergence.lean | 3 +++ src/topology/category/CompHaus/basic.lean | 3 +++ src/topology/category/CompHaus/projective.lean | 3 +++ src/topology/category/Top/limits/basic.lean | 3 +++ src/topology/category/Top/opens.lean | 3 +++ 25 files changed, 75 insertions(+) diff --git a/src/algebraic_topology/nerve.lean b/src/algebraic_topology/nerve.lean index 035a70063220b..09c8373d1a4ba 100644 --- a/src/algebraic_topology/nerve.lean +++ b/src/algebraic_topology/nerve.lean @@ -10,6 +10,9 @@ import algebraic_topology.simplicial_set # The nerve of a category +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file provides the definition of the nerve of a category `C`, which is a simplicial set `nerve C` (see [goerss-jardine-2009], Example I.1.4). diff --git a/src/algebraic_topology/simplicial_set.lean b/src/algebraic_topology/simplicial_set.lean index 48a2fe3bac334..1b42fcba597d4 100644 --- a/src/algebraic_topology/simplicial_set.lean +++ b/src/algebraic_topology/simplicial_set.lean @@ -24,6 +24,9 @@ and their boundaries `∂Δ[n]` and horns `Λ[n, i]`. ## Future work +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + There isn't yet a complete API for simplices, boundaries, and horns. As an example, we should have a function that constructs from a non-surjective order preserving function `fin n → fin n` diff --git a/src/analysis/convex/stone_separation.lean b/src/analysis/convex/stone_separation.lean index b60920f402d7f..903563a4ec4c8 100644 --- a/src/analysis/convex/stone_separation.lean +++ b/src/analysis/convex/stone_separation.lean @@ -8,6 +8,9 @@ import analysis.convex.join /-! # Stone's separation theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file prove Stone's separation theorem. This tells us that any two disjoint convex sets can be separated by a convex set whose complement is also convex. diff --git a/src/category_theory/action.lean b/src/category_theory/action.lean index 2c40856dcc83c..935da21880345 100644 --- a/src/category_theory/action.lean +++ b/src/category_theory/action.lean @@ -12,6 +12,9 @@ import group_theory.semidirect_product /-! # Actions as functors and as categories +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + From a multiplicative action M ↻ X, we can construct a functor from M to the category of types, mapping the single object of M to X and an element `m : M` to map `X → X` given by multiplication by `m`. diff --git a/src/category_theory/limits/filtered_colimit_commutes_finite_limit.lean b/src/category_theory/limits/filtered_colimit_commutes_finite_limit.lean index 4a68a0a2bda04..d320a23a66c89 100644 --- a/src/category_theory/limits/filtered_colimit_commutes_finite_limit.lean +++ b/src/category_theory/limits/filtered_colimit_commutes_finite_limit.lean @@ -13,6 +13,9 @@ import category_theory.concrete_category.basic /-! # Filtered colimits commute with finite limits. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We show that for a functor `F : J × K ⥤ Type v`, when `J` is finite and `K` is filtered, the universal morphism `colimit_limit_to_limit_colimit F` comparing the colimit (over `K`) of the limits (over `J`) with the limit of the colimits is an isomorphism. diff --git a/src/category_theory/noetherian.lean b/src/category_theory/noetherian.lean index 629df87cea753..a622eb69677e4 100644 --- a/src/category_theory/noetherian.lean +++ b/src/category_theory/noetherian.lean @@ -10,6 +10,9 @@ import category_theory.simple /-! # Artinian and noetherian categories +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + An artinian category is a category in which objects do not have infinite decreasing sequences of subobjects. diff --git a/src/category_theory/preadditive/projective.lean b/src/category_theory/preadditive/projective.lean index 6e9bbbfe779f6..9c6098ab570f8 100644 --- a/src/category_theory/preadditive/projective.lean +++ b/src/category_theory/preadditive/projective.lean @@ -11,6 +11,9 @@ import category_theory.limits.preserves.finite /-! # Projective objects and categories with enough projectives +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + An object `P` is called projective if every morphism out of `P` factors through every epimorphism. A category `C` has enough projectives if every object admits an epimorphism from some diff --git a/src/category_theory/simple.lean b/src/category_theory/simple.lean index 55e1e1d62b5b7..1b6d2ee0c5cb5 100644 --- a/src/category_theory/simple.lean +++ b/src/category_theory/simple.lean @@ -12,6 +12,9 @@ import order.atoms /-! # Simple objects +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define simple objects in any category with zero morphisms. A simple object is an object `Y` such that any monomorphism `f : X ⟶ Y` is either an isomorphism or zero (but not both). diff --git a/src/category_theory/sites/adjunction.lean b/src/category_theory/sites/adjunction.lean index 0556200301d8d..d560f26d1e707 100644 --- a/src/category_theory/sites/adjunction.lean +++ b/src/category_theory/sites/adjunction.lean @@ -8,6 +8,9 @@ import category_theory.sites.sheafification import category_theory.sites.whiskering /-! +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we show that an adjunction `F ⊣ G` induces an adjunction between categories of sheaves, under certain hypotheses on `F` and `G`. diff --git a/src/linear_algebra/free_module/finite/matrix.lean b/src/linear_algebra/free_module/finite/matrix.lean index 4f134820b27f9..330c3b24583d0 100644 --- a/src/linear_algebra/free_module/finite/matrix.lean +++ b/src/linear_algebra/free_module/finite/matrix.lean @@ -11,6 +11,9 @@ import linear_algebra.matrix.to_lin /-! # Finite and free modules using matrices +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We provide some instances for finite and free modules involving matrices. ## Main results diff --git a/src/linear_algebra/matrix/basis.lean b/src/linear_algebra/matrix/basis.lean index 2872c03cf147b..972b50dded8c5 100644 --- a/src/linear_algebra/matrix/basis.lean +++ b/src/linear_algebra/matrix/basis.lean @@ -9,6 +9,9 @@ import linear_algebra.matrix.to_lin /-! # Bases and matrices +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the map `basis.to_matrix` that sends a family of vectors to the matrix of their coordinates with respect to some basis. diff --git a/src/linear_algebra/matrix/block.lean b/src/linear_algebra/matrix/block.lean index 0ce6585c06e1f..0b79438b9f1e0 100644 --- a/src/linear_algebra/matrix/block.lean +++ b/src/linear_algebra/matrix/block.lean @@ -10,6 +10,9 @@ import tactic.fin_cases /-! # Block matrices and their determinant +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines a predicate `matrix.block_triangular` saying a matrix is block triangular, and proves the value of the determinant for various matrices built out of blocks. diff --git a/src/linear_algebra/matrix/diagonal.lean b/src/linear_algebra/matrix/diagonal.lean index 24af38dbb344e..107f7f51f388f 100644 --- a/src/linear_algebra/matrix/diagonal.lean +++ b/src/linear_algebra/matrix/diagonal.lean @@ -9,6 +9,9 @@ import linear_algebra.free_module.rank /-! # Diagonal matrices +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains some results on the linear map corresponding to a diagonal matrix (`range`, `ker` and `rank`). diff --git a/src/linear_algebra/matrix/finite_dimensional.lean b/src/linear_algebra/matrix/finite_dimensional.lean index 5931077a135e0..1930c3fab027c 100644 --- a/src/linear_algebra/matrix/finite_dimensional.lean +++ b/src/linear_algebra/matrix/finite_dimensional.lean @@ -11,6 +11,9 @@ import linear_algebra.matrix.to_lin /-! # The finite-dimensional space of matrices +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file shows that `m` by `n` matrices form a finite-dimensional space. Note that this is proven more generally elsewhere over modules as `module.finite.matrix`; this file exists only to provide an entry in the instance list for `finite_dimensional`. diff --git a/src/linear_algebra/matrix/invariant_basis_number.lean b/src/linear_algebra/matrix/invariant_basis_number.lean index 533a22e1d01dc..de2c532e57ac4 100644 --- a/src/linear_algebra/matrix/invariant_basis_number.lean +++ b/src/linear_algebra/matrix/invariant_basis_number.lean @@ -8,6 +8,9 @@ import linear_algebra.invariant_basis_number /-! # Invertible matrices over a ring with invariant basis number are square. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ variables {n m : Type*} [fintype n] [decidable_eq n] [fintype m] [decidable_eq m] diff --git a/src/linear_algebra/matrix/to_lin.lean b/src/linear_algebra/matrix/to_lin.lean index 7e5817ef7353a..949323d8d408f 100644 --- a/src/linear_algebra/matrix/to_lin.lean +++ b/src/linear_algebra/matrix/to_lin.lean @@ -13,6 +13,9 @@ import algebra.algebra.subalgebra.tower /-! # Linear maps and matrices +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the maps to send matrices to a linear map, and to send linear maps between modules with a finite bases to matrices. This defines a linear equivalence between linear maps diff --git a/src/linear_algebra/matrix/transvection.lean b/src/linear_algebra/matrix/transvection.lean index e702144be94ef..8c0daac3129b6 100644 --- a/src/linear_algebra/matrix/transvection.lean +++ b/src/linear_algebra/matrix/transvection.lean @@ -12,6 +12,9 @@ import tactic.field_simp /-! # Transvections +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Transvections are matrices of the form `1 + std_basis_matrix i j c`, where `std_basis_matrix i j c` is the basic matrix with a `c` at position `(i, j)`. Multiplying by such a transvection on the left (resp. on the right) amounts to adding `c` times the `j`-th row to to the `i`-th row diff --git a/src/linear_algebra/matrix/zpow.lean b/src/linear_algebra/matrix/zpow.lean index 70ea8d648dac7..fb7ff4cc6d15f 100644 --- a/src/linear_algebra/matrix/zpow.lean +++ b/src/linear_algebra/matrix/zpow.lean @@ -8,6 +8,9 @@ import linear_algebra.matrix.nonsingular_inverse /-! # Integer powers of square matrices +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we define integer power of matrices, relying on the nonsingular inverse definition for negative powers. diff --git a/src/topology/algebra/module/multilinear.lean b/src/topology/algebra/module/multilinear.lean index 5aebcb46546db..efa86ed869ffc 100644 --- a/src/topology/algebra/module/multilinear.lean +++ b/src/topology/algebra/module/multilinear.lean @@ -9,6 +9,9 @@ import linear_algebra.multilinear.basic /-! # Continuous multilinear maps +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define continuous multilinear maps as maps from `Π(i : ι), M₁ i` to `M₂` which are multilinear and continuous, by extending the space of multilinear maps with a continuity assumption. Here, `M₁ i` and `M₂` are modules over a ring `R`, and `ι` is an arbitrary type, and all these diff --git a/src/topology/algebra/module/strong_topology.lean b/src/topology/algebra/module/strong_topology.lean index 9b60004ef1c63..a1ad2e299b21f 100644 --- a/src/topology/algebra/module/strong_topology.lean +++ b/src/topology/algebra/module/strong_topology.lean @@ -8,6 +8,9 @@ import topology.algebra.uniform_convergence /-! # Strong topologies on the space of continuous linear maps +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we define the strong topologies on `E →L[𝕜] F` associated with a family `𝔖 : set (set E)` to be the topology of uniform convergence on the elements of `𝔖` (also called the topology of `𝔖`-convergence). diff --git a/src/topology/algebra/uniform_convergence.lean b/src/topology/algebra/uniform_convergence.lean index fa23721dcb906..ada5314adc63d 100644 --- a/src/topology/algebra/uniform_convergence.lean +++ b/src/topology/algebra/uniform_convergence.lean @@ -10,6 +10,9 @@ import topology.algebra.filter_basis /-! # Algebraic facts about the topology of uniform convergence +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains algebraic compatibility results about the uniform structure of uniform convergence / `𝔖`-convergence. They will mostly be useful for defining strong topologies on the space of continuous linear maps between two topological vector spaces. diff --git a/src/topology/category/CompHaus/basic.lean b/src/topology/category/CompHaus/basic.lean index 287dc5ac0ea7b..ebb5dd0d43957 100644 --- a/src/topology/category/CompHaus/basic.lean +++ b/src/topology/category/CompHaus/basic.lean @@ -13,6 +13,9 @@ import topology.category.Top.limits.basic /-! # The category of Compact Hausdorff Spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We construct the category of compact Hausdorff spaces. The type of compact Hausdorff spaces is denoted `CompHaus`, and it is endowed with a category instance making it a full subcategory of `Top`. diff --git a/src/topology/category/CompHaus/projective.lean b/src/topology/category/CompHaus/projective.lean index 886fb6fc34f7d..7a9b880053992 100644 --- a/src/topology/category/CompHaus/projective.lean +++ b/src/topology/category/CompHaus/projective.lean @@ -11,6 +11,9 @@ import category_theory.preadditive.projective /-! # CompHaus has enough projectives +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we show that `CompHaus` has enough projectives. ## Main results diff --git a/src/topology/category/Top/limits/basic.lean b/src/topology/category/Top/limits/basic.lean index 68650032113db..9b7985310f6da 100644 --- a/src/topology/category/Top/limits/basic.lean +++ b/src/topology/category/Top/limits/basic.lean @@ -9,6 +9,9 @@ import category_theory.limits.concrete_category /-! # The category of topological spaces has all limits and colimits +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Further, these limits and colimits are preserved by the forgetful functor --- that is, the underlying types are just the limits in the category of types. -/ diff --git a/src/topology/category/Top/opens.lean b/src/topology/category/Top/opens.lean index e69e5a9f4e7e5..c0ea142fc31a2 100644 --- a/src/topology/category/Top/opens.lean +++ b/src/topology/category/Top/opens.lean @@ -11,6 +11,9 @@ import topology.sets.opens /-! # The category of open sets in a topological space. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define `to_Top : opens X ⥤ Top` and `map (f : X ⟶ Y) : opens Y ⥤ opens X`, given by taking preimages of open sets. From cf7a7252c1989efe5800e0b3cdfeb4228ac6b40e Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sat, 29 Apr 2023 09:17:57 +0000 Subject: [PATCH 0905/1291] refactor(algebra/quaternion): remove `quaternion.conj` (#18803) Instead we use `star` directly. We keep only the lemmas that refer to quaternion-specific operators, as the rest are already covered by the star API. In practice, this is mostly the lemmas about the real and imaginary parts of quaternions. The `commute_self_conj` lemma has been replaced by a `is_star_normal` instance. --- src/algebra/quaternion.lean | 259 ++++++------------ .../normed_space/quaternion_exponential.lean | 4 +- src/analysis/quaternion.lean | 17 +- .../clifford_algebra/equivs.lean | 16 +- 4 files changed, 105 insertions(+), 191 deletions(-) diff --git a/src/algebra/quaternion.lean b/src/algebra/quaternion.lean index 8cc429e41dca3..c465eb0e0fb68 100644 --- a/src/algebra/quaternion.lean +++ b/src/algebra/quaternion.lean @@ -22,7 +22,7 @@ algebraic structures on `ℍ[R]`. [quaternion algebra](https://en.wikipedia.org/wiki/Quaternion_algebra) with coefficients `a`, `b` * `quaternion R`, `ℍ[R]` : the space of quaternions, a.k.a. `quaternion_algebra R (-1) (-1)`; * `quaternion.norm_sq` : square of the norm of a quaternion; -* `quaternion.conj` : conjugate of a quaternion; +* `quaternion.star_ring` : provides the conjugate of a quaternion as `has_star.star`; We also define the following algebraic structures on `ℍ[R]`: @@ -339,79 +339,48 @@ by rw [← coe_commutes, coe_mul_eq_smul] lemma smul_coe : x • (y : ℍ[R, c₁, c₂]) = ↑(x * y) := by rw [coe_mul, coe_mul_eq_smul] /-- Quaternion conjugate. -/ -def conj : ℍ[R, c₁, c₂] ≃ₗ[R] ℍ[R, c₁, c₂] := -linear_equiv.of_involutive -{ to_fun := λ a, ⟨a.1, -a.2, -a.3, -a.4⟩, - map_add' := λ a b, by ext; simp [neg_add], - map_smul' := λ r a, by ext; simp } $ -λ a, by simp - -@[simp] lemma re_conj : (conj a).re = a.re := rfl -@[simp] lemma im_i_conj : (conj a).im_i = - a.im_i := rfl -@[simp] lemma im_j_conj : (conj a).im_j = - a.im_j := rfl -@[simp] lemma im_k_conj : (conj a).im_k = - a.im_k := rfl -@[simp] lemma im_conj : (conj a).im = - a.im := ext _ _ neg_zero.symm rfl rfl rfl - -@[simp] lemma conj_mk (a₁ a₂ a₃ a₄ : R) : - conj (mk a₁ a₂ a₃ a₄ : ℍ[R, c₁, c₂]) = ⟨a₁, -a₂, -a₃, -a₄⟩ := -rfl - -@[simp] lemma conj_conj : a.conj.conj = a := ext _ _ rfl (neg_neg _) (neg_neg _) (neg_neg _) +instance : has_star ℍ[R, c₁, c₂] := +{ star := λ a, ⟨a.1, -a.2, -a.3, -a.4⟩ } -lemma conj_add : (a + b).conj = a.conj + b.conj := conj.map_add a b +@[simp] lemma re_star : (star a).re = a.re := rfl +@[simp] lemma im_i_star : (star a).im_i = - a.im_i := rfl +@[simp] lemma im_j_star : (star a).im_j = - a.im_j := rfl +@[simp] lemma im_k_star : (star a).im_k = - a.im_k := rfl +@[simp] lemma im_star : (star a).im = - a.im := ext _ _ neg_zero.symm rfl rfl rfl -@[simp] lemma conj_mul : (a * b).conj = b.conj * a.conj := by ext; simp; ring_exp +@[simp] lemma star_mk (a₁ a₂ a₃ a₄ : R) : + star (mk a₁ a₂ a₃ a₄ : ℍ[R, c₁, c₂]) = ⟨a₁, -a₂, -a₃, -a₄⟩ := +rfl instance : star_ring ℍ[R, c₁, c₂] := -{ star := conj, - star_involutive := conj_conj, - star_add := conj_add, - star_mul := conj_mul } - -@[simp] lemma star_def (a : ℍ[R, c₁, c₂]) : star a = conj a := rfl - -lemma conj_conj_mul : (a.conj * b).conj = b.conj * a := star_star_mul _ _ +{ star_involutive := λ x, by simp [has_star.star], + star_add := λ a b, by ext; simp [neg_add], + star_mul := λ a b, by ext; simp; ring_exp } -lemma conj_mul_conj : (a * b.conj).conj = b * a.conj := star_mul_star _ _ +lemma self_add_star' : a + star a = ↑(2 * a.re) := by ext; simp [two_mul] -lemma self_add_conj' : a + a.conj = ↑(2 * a.re) := by ext; simp [two_mul] +lemma self_add_star : a + star a = 2 * a.re := +by simp only [self_add_star', two_mul, coe_add] -lemma self_add_conj : a + a.conj = 2 * a.re := -by simp only [self_add_conj', two_mul, coe_add] +lemma star_add_self' : star a + a = ↑(2 * a.re) := by rw [add_comm, self_add_star'] -lemma conj_add_self' : a.conj + a = ↑(2 * a.re) := by rw [add_comm, self_add_conj'] +lemma star_add_self : star a + a = 2 * a.re := by rw [add_comm, self_add_star] -lemma conj_add_self : a.conj + a = 2 * a.re := by rw [add_comm, self_add_conj] +lemma star_eq_two_re_sub : star a = ↑(2 * a.re) - a := eq_sub_iff_add_eq.2 a.star_add_self' -lemma conj_eq_two_re_sub : a.conj = ↑(2 * a.re) - a := eq_sub_iff_add_eq.2 a.conj_add_self' - -lemma commute_conj_self : commute a.conj a := -begin - rw [a.conj_eq_two_re_sub], +instance : is_star_normal a := ⟨begin + rw [a.star_eq_two_re_sub], exact (coe_commute (2 * a.re) a).sub_left (commute.refl a) -end - -lemma commute_self_conj : commute a a.conj := -a.commute_conj_self.symm +end⟩ -lemma commute_conj_conj {a b : ℍ[R, c₁, c₂]} (h : commute a b) : commute a.conj b.conj := -h.star_star +@[simp, norm_cast] lemma star_coe : star (x : ℍ[R, c₁, c₂]) = x := by ext; simp -@[simp, norm_cast] lemma conj_coe : conj (x : ℍ[R, c₁, c₂]) = x := by ext; simp +@[simp] lemma star_im : star a.im = - a.im := im_star _ -@[simp] lemma conj_im : conj a.im = - a.im := im_conj _ - -@[simp, norm_cast] lemma conj_nat_cast (n : ℕ) : conj (n : ℍ[R, c₁, c₂]) = n := -@star_nat_cast ℍ[R, c₁, c₂] _ _ n -@[simp, norm_cast] lemma conj_int_cast (z : ℤ) : conj (z : ℍ[R, c₁, c₂]) = z := -@star_int_cast ℍ[R, c₁, c₂] _ _ z - -@[simp] lemma conj_smul [monoid S] [distrib_mul_action S R] (s : S) (a : ℍ[R, c₁, c₂]) : - conj (s • a) = s • conj a := +@[simp] lemma star_smul [monoid S] [distrib_mul_action S R] (s : S) (a : ℍ[R, c₁, c₂]) : + star (s • a) = s • star a := ext _ _ rfl (smul_neg _ _).symm (smul_neg _ _).symm (smul_neg _ _).symm -@[simp] lemma conj_one : conj (1 : ℍ[R, c₁, c₂]) = 1 := conj_coe 1 - lemma eq_re_of_eq_coe {a : ℍ[R, c₁, c₂]} {x : R} (h : a = x) : a = a.re := by rw [h, coe_re] @@ -423,42 +392,33 @@ section char_zero variables [no_zero_divisors R] [char_zero R] @[simp] -lemma conj_eq_self {c₁ c₂ : R} {a : ℍ[R, c₁, c₂]} : - conj a = a ↔ a = a.re := +lemma star_eq_self {c₁ c₂ : R} {a : ℍ[R, c₁, c₂]} : + star a = a ↔ a = a.re := by simp [ext_iff, neg_eq_iff_add_eq_zero, add_self_eq_zero] -lemma conj_eq_neg {c₁ c₂ : R} {a : ℍ[R, c₁, c₂]} : - conj a = -a ↔ a.re = 0 := +lemma star_eq_neg {c₁ c₂ : R} {a : ℍ[R, c₁, c₂]} : + star a = -a ↔ a.re = 0 := by simp [ext_iff, eq_neg_iff_add_eq_zero] end char_zero --- Can't use `rw ← conj_eq_self` in the proof without additional assumptions - -lemma conj_mul_eq_coe : conj a * a = (conj a * a).re := by ext; simp; ring_exp - -lemma mul_conj_eq_coe : a * conj a = (a * conj a).re := -by { rw a.commute_self_conj.eq, exact a.conj_mul_eq_coe } +-- Can't use `rw ← star_eq_self` in the proof without additional assumptions -lemma conj_zero : conj (0 : ℍ[R, c₁, c₂]) = 0 := conj.map_zero +lemma star_mul_eq_coe : star a * a = (star a * a).re := by ext; simp; ring_exp -lemma conj_neg : (-a).conj = -a.conj := (conj : ℍ[R, c₁, c₂] ≃ₗ[R] _).map_neg a - -lemma conj_sub : (a - b).conj = a.conj - b.conj := (conj : ℍ[R, c₁, c₂] ≃ₗ[R] _).map_sub a b - - -@[simp] lemma conj_pow (n : ℕ) : (a ^ n).conj = a.conj ^ n := star_pow _ _ +lemma mul_star_eq_coe : a * star a = (a * star a).re := +by { rw ←star_comm_self', exact a.star_mul_eq_coe } open mul_opposite /-- Quaternion conjugate as an `alg_equiv` to the opposite ring. -/ -def conj_ae : ℍ[R, c₁, c₂] ≃ₐ[R] (ℍ[R, c₁, c₂]ᵐᵒᵖ) := -{ to_fun := op ∘ conj, - inv_fun := conj ∘ unop, +def star_ae : ℍ[R, c₁, c₂] ≃ₐ[R] (ℍ[R, c₁, c₂]ᵐᵒᵖ) := +{ to_fun := op ∘ star, + inv_fun := star ∘ unop, map_mul' := λ x y, by simp, commutes' := λ r, by simp, - .. conj.to_add_equiv.trans op_add_equiv } + .. star_add_equiv.trans op_add_equiv } -@[simp] lemma coe_conj_ae : ⇑(conj_ae : ℍ[R, c₁, c₂] ≃ₐ[R] _) = op ∘ conj := rfl +@[simp] lemma coe_star_ae : ⇑(star_ae : ℍ[R, c₁, c₂] ≃ₐ[R] _) = op ∘ star := rfl end quaternion_algebra @@ -634,56 +594,27 @@ quaternion_algebra.rank_eq_four _ _ lemma finrank_eq_four [strong_rank_condition R] : finite_dimensional.finrank R ℍ[R] = 4 := quaternion_algebra.finrank_eq_four _ _ -/-- Quaternion conjugate. -/ -def conj : ℍ[R] ≃ₗ[R] ℍ[R] := quaternion_algebra.conj +@[simp] lemma star_re : (star a).re = a.re := rfl +@[simp] lemma star_im_i : (star a).im_i = - a.im_i := rfl +@[simp] lemma star_im_j : (star a).im_j = - a.im_j := rfl +@[simp] lemma star_im_k : (star a).im_k = - a.im_k := rfl +@[simp] lemma star_im : (star a).im = - a.im := a.im_star -@[simp] lemma conj_re : a.conj.re = a.re := rfl -@[simp] lemma conj_im_i : a.conj.im_i = - a.im_i := rfl -@[simp] lemma conj_im_j : a.conj.im_j = - a.im_j := rfl -@[simp] lemma conj_im_k : a.conj.im_k = - a.im_k := rfl -@[simp] lemma conj_im : a.conj.im = - a.im := a.im_conj +lemma self_add_star' : a + star a = ↑(2 * a.re) := a.self_add_star' -@[simp] lemma conj_conj : a.conj.conj = a := a.conj_conj +lemma self_add_star : a + star a = 2 * a.re := a.self_add_star -@[simp] lemma conj_add : (a + b).conj = a.conj + b.conj := a.conj_add b +lemma star_add_self' : star a + a = ↑(2 * a.re) := a.star_add_self' -@[simp] lemma conj_mul : (a * b).conj = b.conj * a.conj := a.conj_mul b +lemma star_add_self : star a + a = 2 * a.re := a.star_add_self -lemma conj_conj_mul : (a.conj * b).conj = b.conj * a := a.conj_conj_mul b +lemma star_eq_two_re_sub : star a = ↑(2 * a.re) - a := a.star_eq_two_re_sub -lemma conj_mul_conj : (a * b.conj).conj = b * a.conj := a.conj_mul_conj b +@[simp, norm_cast] lemma star_coe : star (x : ℍ[R]) = x := quaternion_algebra.star_coe x +@[simp] lemma im_star : star a.im = - a.im := quaternion_algebra.im_star _ -lemma self_add_conj' : a + a.conj = ↑(2 * a.re) := a.self_add_conj' - -lemma self_add_conj : a + a.conj = 2 * a.re := a.self_add_conj - -lemma conj_add_self' : a.conj + a = ↑(2 * a.re) := a.conj_add_self' - -lemma conj_add_self : a.conj + a = 2 * a.re := a.conj_add_self - -lemma conj_eq_two_re_sub : a.conj = ↑(2 * a.re) - a := a.conj_eq_two_re_sub - -lemma commute_conj_self : commute a.conj a := a.commute_conj_self - -lemma commute_self_conj : commute a a.conj := a.commute_self_conj - -lemma commute_conj_conj {a b : ℍ[R]} (h : commute a b) : commute a.conj b.conj := -quaternion_algebra.commute_conj_conj h - -alias commute_conj_conj ← commute.quaternion_conj - -@[simp, norm_cast] lemma conj_coe : conj (x : ℍ[R]) = x := quaternion_algebra.conj_coe x -@[simp] lemma im_conj : a.im.conj = - a.im := quaternion_algebra.im_conj _ - -@[simp, norm_cast] lemma conj_nat_cast (n : ℕ) : conj (n : ℍ[R]) = n := -quaternion_algebra.conj_nat_cast _ -@[simp, norm_cast] lemma conj_int_cast (z : ℤ) : conj (z : ℍ[R]) = z := -quaternion_algebra.conj_int_cast _ - -@[simp] lemma conj_smul [monoid S] [distrib_mul_action S R] (s : S) (a : ℍ[R]) : - conj (s • a) = s • conj a := quaternion_algebra.conj_smul _ _ - -@[simp] lemma conj_one : conj (1 : ℍ[R]) = 1 := conj_coe 1 +@[simp] lemma star_smul [monoid S] [distrib_mul_action S R] (s : S) (a : ℍ[R]) : + star (s • a) = s • star a := quaternion_algebra.star_smul _ _ lemma eq_re_of_eq_coe {a : ℍ[R]} {x : R} (h : a = x) : a = a.re := quaternion_algebra.eq_re_of_eq_coe h @@ -694,49 +625,41 @@ quaternion_algebra.eq_re_iff_mem_range_coe section char_zero variables [no_zero_divisors R] [char_zero R] -@[simp] lemma conj_eq_self {a : ℍ[R]} : conj a = a ↔ a = a.re := quaternion_algebra.conj_eq_self +@[simp] lemma star_eq_self {a : ℍ[R]} : star a = a ↔ a = a.re := quaternion_algebra.star_eq_self -@[simp] lemma conj_eq_neg {a : ℍ[R]} : conj a = -a ↔ a.re = 0 := quaternion_algebra.conj_eq_neg +@[simp] lemma star_eq_neg {a : ℍ[R]} : star a = -a ↔ a.re = 0 := quaternion_algebra.star_eq_neg end char_zero -lemma conj_mul_eq_coe : conj a * a = (conj a * a).re := a.conj_mul_eq_coe - -lemma mul_conj_eq_coe : a * conj a = (a * conj a).re := a.mul_conj_eq_coe +lemma star_mul_eq_coe : star a * a = (star a * a).re := a.star_mul_eq_coe -@[simp] lemma conj_zero : conj (0:ℍ[R]) = 0 := quaternion_algebra.conj_zero - -@[simp] lemma conj_neg : (-a).conj = -a.conj := a.conj_neg - -@[simp] lemma conj_sub : (a - b).conj = a.conj - b.conj := a.conj_sub b - -@[simp] lemma conj_pow (n : ℕ) : conj (a ^ n) = conj a ^ n := a.conj_pow n +lemma mul_star_eq_coe : a * star a = (a * star a).re := a.mul_star_eq_coe open mul_opposite /-- Quaternion conjugate as an `alg_equiv` to the opposite ring. -/ -def conj_ae : ℍ[R] ≃ₐ[R] (ℍ[R]ᵐᵒᵖ) := quaternion_algebra.conj_ae +def star_ae : ℍ[R] ≃ₐ[R] (ℍ[R]ᵐᵒᵖ) := quaternion_algebra.star_ae -@[simp] lemma coe_conj_ae : ⇑(conj_ae : ℍ[R] ≃ₐ[R] ℍ[R]ᵐᵒᵖ) = op ∘ conj := rfl +@[simp] lemma coe_star_ae : ⇑(star_ae : ℍ[R] ≃ₐ[R] ℍ[R]ᵐᵒᵖ) = op ∘ star := rfl /-- Square of the norm. -/ def norm_sq : ℍ[R] →*₀ R := -{ to_fun := λ a, (a * a.conj).re, - map_zero' := by rw [conj_zero, zero_mul, zero_re], - map_one' := by rw [conj_one, one_mul, one_re], - map_mul' := λ x y, coe_injective $ by conv_lhs { rw [← mul_conj_eq_coe, conj_mul, mul_assoc, - ← mul_assoc y, y.mul_conj_eq_coe, coe_commutes, ← mul_assoc, x.mul_conj_eq_coe, ← coe_mul] } } +{ to_fun := λ a, (a * star a).re, + map_zero' := by rw [star_zero, zero_mul, zero_re], + map_one' := by rw [star_one, one_mul, one_re], + map_mul' := λ x y, coe_injective $ by conv_lhs { rw [← mul_star_eq_coe, star_mul, mul_assoc, + ← mul_assoc y, y.mul_star_eq_coe, coe_commutes, ← mul_assoc, x.mul_star_eq_coe, ← coe_mul] } } -lemma norm_sq_def : norm_sq a = (a * a.conj).re := rfl +lemma norm_sq_def : norm_sq a = (a * star a).re := rfl lemma norm_sq_def' : norm_sq a = a.1^2 + a.2^2 + a.3^2 + a.4^2 := by simp only [norm_sq_def, sq, mul_neg, sub_neg_eq_add, - mul_re, conj_re, conj_im_i, conj_im_j, conj_im_k] + mul_re, star_re, star_im_i, star_im_j, star_im_k] lemma norm_sq_coe : norm_sq (x : ℍ[R]) = x^2 := -by rw [norm_sq_def, conj_coe, ← coe_mul, coe_re, sq] +by rw [norm_sq_def, star_coe, ← coe_mul, coe_re, sq] -@[simp] lemma norm_sq_conj : norm_sq (conj a) = norm_sq a := by simp [norm_sq_def'] +@[simp] lemma norm_sq_star : norm_sq (star a) = norm_sq a := by simp [norm_sq_def'] @[norm_cast] lemma norm_sq_nat_cast (n : ℕ) : norm_sq (n : ℍ[R]) = n^2 := by rw [←coe_nat_cast, norm_sq_coe] @@ -745,28 +668,28 @@ by rw [←coe_nat_cast, norm_sq_coe] by rw [←coe_int_cast, norm_sq_coe] @[simp] lemma norm_sq_neg : norm_sq (-a) = norm_sq a := -by simp only [norm_sq_def, conj_neg, neg_mul_neg] +by simp only [norm_sq_def, star_neg, neg_mul_neg] -lemma self_mul_conj : a * a.conj = norm_sq a := by rw [mul_conj_eq_coe, norm_sq_def] +lemma self_mul_star : a * star a = norm_sq a := by rw [mul_star_eq_coe, norm_sq_def] -lemma conj_mul_self : a.conj * a = norm_sq a := by rw [← a.commute_self_conj.eq, self_mul_conj] +lemma star_mul_self : star a * a = norm_sq a := by rw [star_comm_self', self_mul_star] lemma im_sq : a.im^2 = -norm_sq a.im := -by simp_rw [sq, ←conj_mul_self, im_conj, neg_mul, neg_neg] +by simp_rw [sq, ←star_mul_self, im_star, neg_mul, neg_neg] lemma coe_norm_sq_add : - (norm_sq (a + b) : ℍ[R]) = norm_sq a + a * b.conj + b * a.conj + norm_sq b := -by simp [← self_mul_conj, mul_add, add_mul, add_assoc] + (norm_sq (a + b) : ℍ[R]) = norm_sq a + a * star b + b * star a + norm_sq b := +by simp [← self_mul_star, mul_add, add_mul, add_assoc] lemma norm_sq_smul (r : R) (q : ℍ[R]) : norm_sq (r • q) = r^2 * norm_sq q := -by simp_rw [norm_sq_def, conj_smul, smul_mul_smul, smul_re, sq, smul_eq_mul] +by simp_rw [norm_sq_def, star_smul, smul_mul_smul, smul_re, sq, smul_eq_mul] -lemma norm_sq_add (a b : ℍ[R]) : norm_sq (a + b) = norm_sq a + norm_sq b + 2 * (a * conj b).re := -calc norm_sq (a + b) = (norm_sq a + (a * conj b).re) + ((b * conj a).re + norm_sq b) - : by simp_rw [norm_sq_def, conj_add, add_mul, mul_add, add_re] - ... = norm_sq a + norm_sq b + ((a * conj b).re + (b * conj a).re) : by abel - ... = norm_sq a + norm_sq b + 2 * (a * conj b).re - : by rw [←add_re, ←conj_mul_conj a b, self_add_conj', coe_re] +lemma norm_sq_add (a b : ℍ[R]) : norm_sq (a + b) = norm_sq a + norm_sq b + 2 * (a * star b).re := +calc norm_sq (a + b) = (norm_sq a + (a * star b).re) + ((b * star a).re + norm_sq b) + : by simp_rw [norm_sq_def, star_add, add_mul, mul_add, add_re] + ... = norm_sq a + norm_sq b + ((a * star b).re + (b * star a).re) : by abel + ... = norm_sq a + norm_sq b + 2 * (a * star b).re + : by rw [←add_re, ←star_mul_star a b, self_add_star', coe_re] end quaternion @@ -808,18 +731,18 @@ no_zero_divisors.to_is_domain _ lemma sq_eq_norm_sq : a^2 = norm_sq a ↔ a = a.re := begin - simp_rw [←conj_eq_self], + simp_rw [←star_eq_self], obtain rfl | hq0 := eq_or_ne a 0, { simp }, - { rw [←conj_mul_self, sq, mul_left_inj' hq0, eq_comm] } + { rw [←star_mul_self, sq, mul_left_inj' hq0, eq_comm] } end lemma sq_eq_neg_norm_sq : a^2 = -norm_sq a ↔ a.re = 0 := begin - simp_rw [←conj_eq_neg], + simp_rw [←star_eq_neg], obtain rfl | hq0 := eq_or_ne a 0, { simp }, - rw [←conj_mul_self, ←mul_neg, ←neg_sq, sq, mul_left_inj' (neg_ne_zero.mpr hq0), eq_comm], + rw [←star_mul_self, ←mul_neg, ←neg_sq, sq, mul_left_inj' (neg_ne_zero.mpr hq0), eq_comm], end end linear_ordered_comm_ring @@ -828,12 +751,12 @@ section field variables [linear_ordered_field R] (a b : ℍ[R]) -@[simps { attrs := [] }] instance : has_inv ℍ[R] := ⟨λ a, (norm_sq a)⁻¹ • a.conj⟩ +@[simps { attrs := [] }] instance : has_inv ℍ[R] := ⟨λ a, (norm_sq a)⁻¹ • star a⟩ instance : group_with_zero ℍ[R] := { inv := has_inv.inv, - inv_zero := by rw [has_inv_inv, conj_zero, smul_zero], - mul_inv_cancel := λ a ha, by rw [has_inv_inv, algebra.mul_smul_comm, self_mul_conj, smul_coe, + inv_zero := by rw [has_inv_inv, star_zero, smul_zero], + mul_inv_cancel := λ a ha, by rw [has_inv_inv, algebra.mul_smul_comm, self_mul_star, smul_coe, inv_mul_cancel (norm_sq_ne_zero.2 ha), coe_one], .. quaternion.nontrivial, .. (by apply_instance : monoid_with_zero ℍ[R]) } @@ -865,10 +788,6 @@ instance : division_ring ℍ[R] := @[simp, norm_cast] lemma rat_cast_im (q : ℚ) : (q : ℍ[R]).im = 0 := rfl @[norm_cast] lemma coe_rat_cast (q : ℚ) : ↑(q : R) = (q : ℍ[R]) := rfl -lemma conj_inv : conj (a⁻¹) = (conj a)⁻¹ := star_inv' a -lemma conj_zpow (z : ℤ) : conj (a ^ z) = conj a ^ z := star_zpow₀ a z -@[simp, norm_cast] lemma conj_rat_cast (q : ℚ) : conj (q : ℍ[R]) = q := @star_rat_cast ℍ[R] _ _ q - @[simp] lemma norm_sq_inv : norm_sq a⁻¹ = (norm_sq a)⁻¹ := map_inv₀ norm_sq _ @[simp] lemma norm_sq_div : norm_sq (a / b) = norm_sq a / norm_sq b := map_div₀ norm_sq a b @[simp] lemma norm_sq_zpow (z : ℤ) : norm_sq (a ^ z) = norm_sq a ^ z := map_zpow₀ norm_sq a z diff --git a/src/analysis/normed_space/quaternion_exponential.lean b/src/analysis/normed_space/quaternion_exponential.lean index 98345379d5517..dec605b571e51 100644 --- a/src/analysis/normed_space/quaternion_exponential.lean +++ b/src/analysis/normed_space/quaternion_exponential.lean @@ -27,8 +27,6 @@ open_locale quaternion nat namespace quaternion -lemma conj_exp (q : ℍ[ℝ]) : conj (exp ℝ q) = exp ℝ (conj q) := star_exp q - @[simp, norm_cast] lemma exp_coe (r : ℝ) : exp ℝ (r : ℍ[ℝ]) = ↑(exp ℝ r) := (map_exp ℝ (algebra_map ℝ ℍ[ℝ]) (continuous_algebra_map _ _) _).symm @@ -119,7 +117,7 @@ calc norm_sq (exp ℝ q) congr' 1, obtain hv | hv := eq_or_ne (‖q.im‖) 0, { simp [hv] }, - rw [norm_sq_add, norm_sq_smul, conj_smul, coe_mul_eq_smul, smul_re, smul_re, conj_re, im_re, + rw [norm_sq_add, norm_sq_smul, star_smul, coe_mul_eq_smul, smul_re, smul_re, star_re, im_re, smul_zero, smul_zero, mul_zero, add_zero, div_pow, norm_sq_coe, norm_sq_eq_norm_sq, ←sq, div_mul_cancel _ (pow_ne_zero _ hv)], end diff --git a/src/analysis/quaternion.lean b/src/analysis/quaternion.lean index 02fe08b71b90e..35bf8ff9e3bd3 100644 --- a/src/analysis/quaternion.lean +++ b/src/analysis/quaternion.lean @@ -35,11 +35,11 @@ open_locale real_inner_product_space namespace quaternion -instance : has_inner ℝ ℍ := ⟨λ a b, (a * b.conj).re⟩ +instance : has_inner ℝ ℍ := ⟨λ a b, (a * star b).re⟩ lemma inner_self (a : ℍ) : ⟪a, a⟫ = norm_sq a := rfl -lemma inner_def (a b : ℍ) : ⟪a, b⟫ = (a * b.conj).re := rfl +lemma inner_def (a b : ℍ) : ⟪a, b⟫ = (a * star b).re := rfl noncomputable instance : normed_add_comm_group ℍ := @inner_product_space.of_core.to_normed_add_comm_group ℝ ℍ _ _ _ @@ -65,11 +65,11 @@ by rw [norm_eq_sqrt_real_inner, inner_self, norm_sq_coe, real.sqrt_sq_eq_abs, re @[simp, norm_cast] lemma nnnorm_coe (a : ℝ) : ‖(a : ℍ)‖₊ = ‖a‖₊ := subtype.ext $ norm_coe a -@[simp] lemma norm_conj (a : ℍ) : ‖conj a‖ = ‖a‖ := -by simp_rw [norm_eq_sqrt_real_inner, inner_self, norm_sq_conj] +@[simp] lemma norm_star (a : ℍ) : ‖star a‖ = ‖a‖ := +by simp_rw [norm_eq_sqrt_real_inner, inner_self, norm_sq_star] -@[simp] lemma nnnorm_conj (a : ℍ) : ‖conj a‖₊ = ‖a‖₊ := -subtype.ext $ norm_conj a +@[simp] lemma nnnorm_star (a : ℍ) : ‖star a‖₊ = ‖a‖₊ := +subtype.ext $ norm_star a noncomputable instance : normed_division_ring ℍ := { dist_eq := λ _ _, rfl, @@ -81,7 +81,7 @@ instance : normed_algebra ℝ ℍ := to_algebra := (quaternion.algebra : algebra ℝ ℍ) } instance : cstar_ring ℍ := -{ norm_star_mul_self := λ x, (norm_mul _ _).trans $ congr_arg (* ‖x‖) (norm_conj x) } +{ norm_star_mul_self := λ x, (norm_mul _ _).trans $ congr_arg (* ‖x‖) (norm_star x) } instance : has_coe ℂ ℍ := ⟨λ z, ⟨z.re, z.im, 0, 0⟩⟩ @@ -128,9 +128,6 @@ noncomputable def linear_isometry_equiv_tuple : ℍ ≃ₗᵢ[ℝ] euclidean_spa ..(quaternion_algebra.linear_equiv_tuple (-1 : ℝ) (-1 : ℝ)).trans (pi_Lp.linear_equiv 2 ℝ (λ _ : fin 4, ℝ)).symm } -@[continuity] lemma continuous_conj : continuous (conj : ℍ → ℍ) := -continuous_star - @[continuity] lemma continuous_coe : continuous (coe : ℝ → ℍ) := continuous_algebra_map ℝ ℍ diff --git a/src/linear_algebra/clifford_algebra/equivs.lean b/src/linear_algebra/clifford_algebra/equivs.lean index a7a3bbbcb4dc3..fbfc724ea7963 100644 --- a/src/linear_algebra/clifford_algebra/equivs.lean +++ b/src/linear_algebra/clifford_algebra/equivs.lean @@ -48,7 +48,7 @@ We show additionally that this equivalence sends `quaternion_algebra.conj` to th and vice-versa: * `clifford_algebra_quaternion.to_quaternion_star` -* `clifford_algebra_quaternion.of_quaternion_conj` +* `clifford_algebra_quaternion.of_quaternion_star` ## Dual numbers @@ -273,20 +273,20 @@ clifford_algebra.lift_ι_apply _ _ v /-- The "clifford conjugate" maps to the quaternion conjugate. -/ lemma to_quaternion_star (c : clifford_algebra (Q c₁ c₂)) : - to_quaternion (star c) = quaternion_algebra.conj (to_quaternion c) := + to_quaternion (star c) = star (to_quaternion c) := begin simp only [clifford_algebra.star_def'], induction c using clifford_algebra.induction, case h_grade0 : r { simp only [reverse.commutes, alg_hom.commutes, quaternion_algebra.coe_algebra_map, - quaternion_algebra.conj_coe], }, + quaternion_algebra.star_coe], }, case h_grade1 : x { rw [reverse_ι, involute_ι, to_quaternion_ι, alg_hom.map_neg, to_quaternion_ι, - quaternion_algebra.neg_mk, conj_mk, neg_zero], }, + quaternion_algebra.neg_mk, star_mk, neg_zero], }, case h_mul : x₁ x₂ hx₁ hx₂ - { simp only [reverse.map_mul, alg_hom.map_mul, hx₁, hx₂, quaternion_algebra.conj_mul] }, + { simp only [reverse.map_mul, alg_hom.map_mul, hx₁, hx₂, star_mul] }, case h_add : x₁ x₂ hx₁ hx₂ - { simp only [reverse.map_add, alg_hom.map_add, hx₁, hx₂, quaternion_algebra.conj_add] }, + { simp only [reverse.map_add, alg_hom.map_add, hx₁, hx₂, star_add] }, end /-- Map a quaternion into the clifford algebra. -/ @@ -339,8 +339,8 @@ alg_equiv.of_alg_hom to_quaternion of_quaternion of_quaternion_comp_to_quaternion /-- The quaternion conjugate maps to the "clifford conjugate" (aka `star`). -/ -@[simp] lemma of_quaternion_conj (q : ℍ[R,c₁,c₂]) : - of_quaternion (q.conj) = star (of_quaternion q) := +@[simp] lemma of_quaternion_star (q : ℍ[R,c₁,c₂]) : + of_quaternion (star q) = star (of_quaternion q) := clifford_algebra_quaternion.equiv.injective $ by rw [equiv_apply, equiv_apply, to_quaternion_star, to_quaternion_of_quaternion, to_quaternion_of_quaternion] From 7cc39182f1aa354c754bab390c9ff9bef6caffef Mon Sep 17 00:00:00 2001 From: Alex J Best Date: Sat, 29 Apr 2023 10:53:16 +0000 Subject: [PATCH 0906/1291] chore(.github/workflows): update checkout and setup python versions (#18764) To stay ahead of the node 12 to node 16 deprecation due "summer 2023" https://github.blog/changelog/2022-09-22-github-actions-all-actions-will-begin-running-on-node16-instead-of-node12/ --- .github/workflows/add_port_comment.yml | 2 +- .github/workflows/add_ported_warnings.yml | 2 +- .github/workflows/bors.yml | 12 ++++++------ .github/workflows/build.yml | 12 ++++++------ .github/workflows/build.yml.in | 12 ++++++------ .github/workflows/build_fork.yml | 12 ++++++------ .github/workflows/lint_self_test.yml | 4 ++-- 7 files changed, 28 insertions(+), 28 deletions(-) diff --git a/.github/workflows/add_port_comment.yml b/.github/workflows/add_port_comment.yml index a65ed98ec6892..d00eabc7f42d0 100644 --- a/.github/workflows/add_port_comment.yml +++ b/.github/workflows/add_port_comment.yml @@ -13,7 +13,7 @@ jobs: - uses: actions/checkout@v3 - name: install Python - uses: actions/setup-python@v3 + uses: actions/setup-python@v4 with: python-version: 3.8 diff --git a/.github/workflows/add_ported_warnings.yml b/.github/workflows/add_ported_warnings.yml index 7530aa696e99c..767ac66b05a19 100644 --- a/.github/workflows/add_ported_warnings.yml +++ b/.github/workflows/add_ported_warnings.yml @@ -11,7 +11,7 @@ jobs: - uses: actions/checkout@v3 - name: install Python - uses: actions/setup-python@v3 + uses: actions/setup-python@v4 with: python-version: 3.8 diff --git a/.github/workflows/bors.yml b/.github/workflows/bors.yml index b35eb1d373d3d..6db50bfd9271b 100644 --- a/.github/workflows/bors.yml +++ b/.github/workflows/bors.yml @@ -32,7 +32,7 @@ jobs: name: Lint style runs-on: bors steps: - - uses: actions/checkout@v2 + - uses: actions/checkout@v3 - name: Install bibtool if: ${{ 'bors' == 'ubuntu-latest' }} @@ -42,7 +42,7 @@ jobs: - name: install Python if: ${{ 'bors' == 'ubuntu-latest' }} - uses: actions/setup-python@v2 + uses: actions/setup-python@v4 with: python-version: 3.8 @@ -64,7 +64,7 @@ jobs: outputs: artifact_name: ${{ steps.setup_precompiled.outputs.artifact_name }} steps: - - uses: actions/checkout@v2 + - uses: actions/checkout@v3 with: fetch-depth: ${{ env.GIT_HISTORY_DEPTH }} @@ -78,7 +78,7 @@ jobs: - name: install Python if: ${{ ! 1 }} - uses: actions/setup-python@v1 + uses: actions/setup-python@v4 with: python-version: 3.8 @@ -196,7 +196,7 @@ jobs: - name: install Python if: ${{ ! 1 }} - uses: actions/setup-python@v1 + uses: actions/setup-python@v4 with: python-version: 3.8 @@ -257,7 +257,7 @@ jobs: needs: [style_lint, build, lint, tests] runs-on: ubuntu-latest steps: - - uses: actions/checkout@v2 + - uses: actions/checkout@v3 - id: PR uses: 8BitJonny/gh-get-current-pr@1.2.0 diff --git a/.github/workflows/build.yml b/.github/workflows/build.yml index f39161df1a753..91486bf89a0f8 100644 --- a/.github/workflows/build.yml +++ b/.github/workflows/build.yml @@ -40,7 +40,7 @@ jobs: name: Lint style runs-on: ubuntu-latest steps: - - uses: actions/checkout@v2 + - uses: actions/checkout@v3 - name: Install bibtool if: ${{ 'ubuntu-latest' == 'ubuntu-latest' }} @@ -50,7 +50,7 @@ jobs: - name: install Python if: ${{ 'ubuntu-latest' == 'ubuntu-latest' }} - uses: actions/setup-python@v2 + uses: actions/setup-python@v4 with: python-version: 3.8 @@ -72,7 +72,7 @@ jobs: outputs: artifact_name: ${{ steps.setup_precompiled.outputs.artifact_name }} steps: - - uses: actions/checkout@v2 + - uses: actions/checkout@v3 with: fetch-depth: ${{ env.GIT_HISTORY_DEPTH }} @@ -86,7 +86,7 @@ jobs: - name: install Python if: ${{ ! 1 }} - uses: actions/setup-python@v1 + uses: actions/setup-python@v4 with: python-version: 3.8 @@ -204,7 +204,7 @@ jobs: - name: install Python if: ${{ ! 1 }} - uses: actions/setup-python@v1 + uses: actions/setup-python@v4 with: python-version: 3.8 @@ -265,7 +265,7 @@ jobs: needs: [style_lint, build, lint, tests] runs-on: ubuntu-latest steps: - - uses: actions/checkout@v2 + - uses: actions/checkout@v3 - id: PR uses: 8BitJonny/gh-get-current-pr@1.2.0 diff --git a/.github/workflows/build.yml.in b/.github/workflows/build.yml.in index d524f35872bb5..8b0fd71807cd4 100644 --- a/.github/workflows/build.yml.in +++ b/.github/workflows/build.yml.in @@ -18,7 +18,7 @@ jobs: name: Lint styleJOB_NAME runs-on: STYLE_LINT_RUNNER steps: - - uses: actions/checkout@v2 + - uses: actions/checkout@v3 - name: Install bibtool if: ${{ 'STYLE_LINT_RUNNER' == 'ubuntu-latest' }} @@ -28,7 +28,7 @@ jobs: - name: install Python if: ${{ 'STYLE_LINT_RUNNER' == 'ubuntu-latest' }} - uses: actions/setup-python@v2 + uses: actions/setup-python@v4 with: python-version: 3.8 @@ -50,7 +50,7 @@ jobs: outputs: artifact_name: ${{ steps.setup_precompiled.outputs.artifact_name }} steps: - - uses: actions/checkout@v2 + - uses: actions/checkout@v3 with: fetch-depth: ${{ env.GIT_HISTORY_DEPTH }} @@ -64,7 +64,7 @@ jobs: - name: install Python if: ${{ ! IS_SELF_HOSTED }} - uses: actions/setup-python@v1 + uses: actions/setup-python@v4 with: python-version: 3.8 @@ -182,7 +182,7 @@ jobs: - name: install Python if: ${{ ! IS_SELF_HOSTED }} - uses: actions/setup-python@v1 + uses: actions/setup-python@v4 with: python-version: 3.8 @@ -243,7 +243,7 @@ jobs: needs: [style_lint, build, lint, tests] runs-on: ubuntu-latest steps: - - uses: actions/checkout@v2 + - uses: actions/checkout@v3 - id: PR uses: 8BitJonny/gh-get-current-pr@1.2.0 diff --git a/.github/workflows/build_fork.yml b/.github/workflows/build_fork.yml index 782a118a88ba0..3d8a9c1c29c17 100644 --- a/.github/workflows/build_fork.yml +++ b/.github/workflows/build_fork.yml @@ -38,7 +38,7 @@ jobs: name: Lint style (fork) runs-on: ubuntu-latest steps: - - uses: actions/checkout@v2 + - uses: actions/checkout@v3 - name: Install bibtool if: ${{ 'ubuntu-latest' == 'ubuntu-latest' }} @@ -48,7 +48,7 @@ jobs: - name: install Python if: ${{ 'ubuntu-latest' == 'ubuntu-latest' }} - uses: actions/setup-python@v2 + uses: actions/setup-python@v4 with: python-version: 3.8 @@ -70,7 +70,7 @@ jobs: outputs: artifact_name: ${{ steps.setup_precompiled.outputs.artifact_name }} steps: - - uses: actions/checkout@v2 + - uses: actions/checkout@v3 with: fetch-depth: ${{ env.GIT_HISTORY_DEPTH }} @@ -84,7 +84,7 @@ jobs: - name: install Python if: ${{ ! 0 }} - uses: actions/setup-python@v1 + uses: actions/setup-python@v4 with: python-version: 3.8 @@ -202,7 +202,7 @@ jobs: - name: install Python if: ${{ ! 0 }} - uses: actions/setup-python@v1 + uses: actions/setup-python@v4 with: python-version: 3.8 @@ -263,7 +263,7 @@ jobs: needs: [style_lint, build, lint, tests] runs-on: ubuntu-latest steps: - - uses: actions/checkout@v2 + - uses: actions/checkout@v3 - id: PR uses: 8BitJonny/gh-get-current-pr@1.2.0 diff --git a/.github/workflows/lint_self_test.yml b/.github/workflows/lint_self_test.yml index fdbe6c5be4988..235bf36b45a96 100644 --- a/.github/workflows/lint_self_test.yml +++ b/.github/workflows/lint_self_test.yml @@ -26,9 +26,9 @@ jobs: - name: 3.8 - name: 3.9 steps: - - uses: actions/checkout@v2 + - uses: actions/checkout@v3 - name: Set up Python ${{ matrix.python-version.name }} - uses: actions/setup-python@v2 + uses: actions/setup-python@v4 with: python-version: ${{ matrix.python-version.name }} From 86d04064ca33ee3d3405fbfc497d494fd2dd4796 Mon Sep 17 00:00:00 2001 From: Alex J Best Date: Sat, 29 Apr 2023 12:39:05 +0000 Subject: [PATCH 0907/1291] chore(.github/workflows): update get-changed-files and download-artifact version (#18765) To stay ahead of the node 12 to node 16 deprecation due "summer 2023" https://github.blog/changelog/2022-09-22-github-actions-all-actions-will-begin-running-on-node16-instead-of-node12/ --- .github/workflows/add_ported_warnings.yml | 2 +- .github/workflows/bors.yml | 4 ++-- .github/workflows/build.yml | 4 ++-- .github/workflows/build.yml.in | 4 ++-- .github/workflows/build_fork.yml | 4 ++-- 5 files changed, 9 insertions(+), 9 deletions(-) diff --git a/.github/workflows/add_ported_warnings.yml b/.github/workflows/add_ported_warnings.yml index 767ac66b05a19..689dd520f3ac4 100644 --- a/.github/workflows/add_ported_warnings.yml +++ b/.github/workflows/add_ported_warnings.yml @@ -22,7 +22,7 @@ jobs: # TODO: is this really faster than just calling git from python? - name: Get changed files id: changed-files - uses: Ana06/get-changed-files@v1.2 + uses: Ana06/get-changed-files@v2.2.0 - name: run the script id: script diff --git a/.github/workflows/bors.yml b/.github/workflows/bors.yml index 6db50bfd9271b..a742b9ae0d2da 100644 --- a/.github/workflows/bors.yml +++ b/.github/workflows/bors.yml @@ -143,7 +143,7 @@ jobs: run: rm -rf ./* ./.??* - name: retrieve build - uses: actions/download-artifact@v2 + uses: actions/download-artifact@v3 with: name: ${{ needs.build.outputs.artifact_name }} @@ -180,7 +180,7 @@ jobs: run: rm -rf ./* ./.??* - name: retrieve build - uses: actions/download-artifact@v2 + uses: actions/download-artifact@v3 with: name: ${{ needs.build.outputs.artifact_name }} diff --git a/.github/workflows/build.yml b/.github/workflows/build.yml index 91486bf89a0f8..0b696f4e052a1 100644 --- a/.github/workflows/build.yml +++ b/.github/workflows/build.yml @@ -151,7 +151,7 @@ jobs: run: rm -rf ./* ./.??* - name: retrieve build - uses: actions/download-artifact@v2 + uses: actions/download-artifact@v3 with: name: ${{ needs.build.outputs.artifact_name }} @@ -188,7 +188,7 @@ jobs: run: rm -rf ./* ./.??* - name: retrieve build - uses: actions/download-artifact@v2 + uses: actions/download-artifact@v3 with: name: ${{ needs.build.outputs.artifact_name }} diff --git a/.github/workflows/build.yml.in b/.github/workflows/build.yml.in index 8b0fd71807cd4..dab918a1a7e7b 100644 --- a/.github/workflows/build.yml.in +++ b/.github/workflows/build.yml.in @@ -129,7 +129,7 @@ jobs: run: rm -rf ./* ./.??* - name: retrieve build - uses: actions/download-artifact@v2 + uses: actions/download-artifact@v3 with: name: ${{ needs.build.outputs.artifact_name }} @@ -166,7 +166,7 @@ jobs: run: rm -rf ./* ./.??* - name: retrieve build - uses: actions/download-artifact@v2 + uses: actions/download-artifact@v3 with: name: ${{ needs.build.outputs.artifact_name }} diff --git a/.github/workflows/build_fork.yml b/.github/workflows/build_fork.yml index 3d8a9c1c29c17..2541ccc498838 100644 --- a/.github/workflows/build_fork.yml +++ b/.github/workflows/build_fork.yml @@ -149,7 +149,7 @@ jobs: run: rm -rf ./* ./.??* - name: retrieve build - uses: actions/download-artifact@v2 + uses: actions/download-artifact@v3 with: name: ${{ needs.build.outputs.artifact_name }} @@ -186,7 +186,7 @@ jobs: run: rm -rf ./* ./.??* - name: retrieve build - uses: actions/download-artifact@v2 + uses: actions/download-artifact@v3 with: name: ${{ needs.build.outputs.artifact_name }} From 011cafb4a5bc695875d186e245d6b3df03bf6c40 Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Sun, 30 Apr 2023 06:51:56 +0000 Subject: [PATCH 0908/1291] chore (measure_theory/integral): split interval_integral.lean into two (#18898) This divides `interval_integral.lean` (currently 2800 lines) into two roughly equal chunks -- moving the variants of FTC into a new file `fund_thm_calculus.lean` and leaving definitions and basic lemmas in the original file. Co-authored-by: sgouezel --- .../100-theorems-list/9_area_of_a_circle.lean | 2 +- src/analysis/fourier/add_circle.lean | 1 + src/analysis/special_functions/integrals.lean | 2 +- .../integral/fund_thm_calculus.lean | 1485 +++++++++++++++++ .../integral/integral_eq_improper.lean | 2 +- .../integral/interval_integral.lean | 1463 +--------------- 6 files changed, 1493 insertions(+), 1462 deletions(-) create mode 100644 src/measure_theory/integral/fund_thm_calculus.lean diff --git a/archive/100-theorems-list/9_area_of_a_circle.lean b/archive/100-theorems-list/9_area_of_a_circle.lean index 5e6177e2c5c98..9b628e2f8cc89 100644 --- a/archive/100-theorems-list/9_area_of_a_circle.lean +++ b/archive/100-theorems-list/9_area_of_a_circle.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 James Arthur, Benjamin Davidson, Andrew Souther. All rights r Released under Apache 2.0 license as described in the file LICENSE. Authors: James Arthur, Benjamin Davidson, Andrew Souther -/ -import measure_theory.integral.interval_integral +import measure_theory.integral.fund_thm_calculus import analysis.special_functions.sqrt import analysis.special_functions.trigonometric.inverse_deriv diff --git a/src/analysis/fourier/add_circle.lean b/src/analysis/fourier/add_circle.lean index da480b55ef038..f96ad82907a0d 100644 --- a/src/analysis/fourier/add_circle.lean +++ b/src/analysis/fourier/add_circle.lean @@ -11,6 +11,7 @@ import measure_theory.function.l2_space import measure_theory.group.integration import measure_theory.integral.periodic import topology.continuous_function.stone_weierstrass +import measure_theory.integral.fund_thm_calculus /-! diff --git a/src/analysis/special_functions/integrals.lean b/src/analysis/special_functions/integrals.lean index b7e245c98d162..9291d26b5273a 100644 --- a/src/analysis/special_functions/integrals.lean +++ b/src/analysis/special_functions/integrals.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Benjamin Davidson. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Benjamin Davidson -/ -import measure_theory.integral.interval_integral +import measure_theory.integral.fund_thm_calculus import analysis.special_functions.trigonometric.arctan_deriv /-! diff --git a/src/measure_theory/integral/fund_thm_calculus.lean b/src/measure_theory/integral/fund_thm_calculus.lean new file mode 100644 index 0000000000000..43b4c4da776a4 --- /dev/null +++ b/src/measure_theory/integral/fund_thm_calculus.lean @@ -0,0 +1,1485 @@ +/- +Copyright (c) 2020 Yury G. Kudryashov. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yury G. Kudryashov, Patrick Massot, Sébastien Gouëzel +-/ +import analysis.calculus.fderiv_measurable +import analysis.normed_space.dual +import measure_theory.integral.interval_integral +import measure_theory.integral.vitali_caratheodory + +/-! +# Fundamental Theorem of Calculus + +We prove various versions of the +[fundamental theorem of calculus](https://en.wikipedia.org/wiki/Fundamental_theorem_of_calculus) +for interval integrals in `ℝ`. + +Recall that its first version states that the function `(u, v) ↦ ∫ x in u..v, f x` has derivative +`(δu, δv) ↦ δv • f b - δu • f a` at `(a, b)` provided that `f` is continuous at `a` and `b`, +and its second version states that, if `f` has an integrable derivative on `[a, b]`, then +`∫ x in a..b, f' x = f b - f a`. + +## Main statements + +### FTC-1 for Lebesgue measure + +We prove several versions of FTC-1, all in the `interval_integral` namespace. Many of them follow +the naming scheme `integral_has(_strict?)_(f?)deriv(_within?)_at(_of_tendsto_ae?)(_right|_left?)`. +They formulate FTC in terms of `has(_strict?)_(f?)deriv(_within?)_at`. +Let us explain the meaning of each part of the name: + +* `_strict` means that the theorem is about strict differentiability; +* `f` means that the theorem is about differentiability in both endpoints; incompatible with + `_right|_left`; +* `_within` means that the theorem is about one-sided derivatives, see below for details; +* `_of_tendsto_ae` means that instead of continuity the theorem assumes that `f` has a finite limit + almost surely as `x` tends to `a` and/or `b`; +* `_right` or `_left` mean that the theorem is about differentiability in the right (resp., left) + endpoint. + +We also reformulate these theorems in terms of `(f?)deriv(_within?)`. These theorems are named +`(f?)deriv(_within?)_integral(_of_tendsto_ae?)(_right|_left?)` with the same meaning of parts of the +name. + +### One-sided derivatives + +Theorem `integral_has_fderiv_within_at_of_tendsto_ae` states that `(u, v) ↦ ∫ x in u..v, f x` has a +derivative `(δu, δv) ↦ δv • cb - δu • ca` within the set `s × t` at `(a, b)` provided that `f` tends +to `ca` (resp., `cb`) almost surely at `la` (resp., `lb`), where possible values of `s`, `t`, and +corresponding filters `la`, `lb` are given in the following table. + +| `s` | `la` | `t` | `lb` | +| ------- | ---- | --- | ---- | +| `Iic a` | `𝓝[≤] a` | `Iic b` | `𝓝[≤] b` | +| `Ici a` | `𝓝[>] a` | `Ici b` | `𝓝[>] b` | +| `{a}` | `⊥` | `{b}` | `⊥` | +| `univ` | `𝓝 a` | `univ` | `𝓝 b` | + +We use a typeclass `FTC_filter` to make Lean automatically find `la`/`lb` based on `s`/`t`. This way +we can formulate one theorem instead of `16` (or `8` if we leave only non-trivial ones not covered +by `integral_has_deriv_within_at_of_tendsto_ae_(left|right)` and +`integral_has_fderiv_at_of_tendsto_ae`). Similarly, +`integral_has_deriv_within_at_of_tendsto_ae_right` works for both one-sided derivatives using the +same typeclass to find an appropriate filter. + +### FTC for a locally finite measure + +Before proving FTC for the Lebesgue measure, we prove a few statements that can be seen as FTC for +any measure. The most general of them, +`measure_integral_sub_integral_sub_linear_is_o_of_tendsto_ae`, states the following. Let `(la, la')` +be an `FTC_filter` pair of filters around `a` (i.e., `FTC_filter a la la'`) and let `(lb, lb')` be +an `FTC_filter` pair of filters around `b`. If `f` has finite limits `ca` and `cb` almost surely at +`la'` and `lb'`, respectively, then +`∫ x in va..vb, f x ∂μ - ∫ x in ua..ub, f x ∂μ = ∫ x in ub..vb, cb ∂μ - ∫ x in ua..va, ca ∂μ + + o(‖∫ x in ua..va, (1:ℝ) ∂μ‖ + ‖∫ x in ub..vb, (1:ℝ) ∂μ‖)` as `ua` and `va` tend to `la` while +`ub` and `vb` tend to `lb`. + +### FTC-2 and corollaries + +We use FTC-1 to prove several versions of FTC-2 for the Lebesgue measure, using a similar naming +scheme as for the versions of FTC-1. They include: +* `interval_integral.integral_eq_sub_of_has_deriv_right_of_le` - most general version, for functions + with a right derivative +* `interval_integral.integral_eq_sub_of_has_deriv_at'` - version for functions with a derivative on + an open set +* `interval_integral.integral_deriv_eq_sub'` - version that is easiest to use when computing the + integral of a specific function + +We then derive additional integration techniques from FTC-2: +* `interval_integral.integral_mul_deriv_eq_deriv_mul` - integration by parts +* `interval_integral.integral_comp_mul_deriv''` - integration by substitution + +Many applications of these theorems can be found in the file `analysis.special_functions.integrals`. + +Note that the assumptions of FTC-2 are formulated in the form that `f'` is integrable. To use it in +a context with the stronger assumption that `f'` is continuous, one can use +`continuous_on.interval_integrable` or `continuous_on.integrable_on_Icc` or +`continuous_on.integrable_on_interval`. + +### `FTC_filter` class + +As explained above, many theorems in this file rely on the typeclass +`FTC_filter (a : ℝ) (l l' : filter ℝ)` to avoid code duplication. This typeclass combines four +assumptions: + +- `pure a ≤ l`; +- `l' ≤ 𝓝 a`; +- `l'` has a basis of measurable sets; +- if `u n` and `v n` tend to `l`, then for any `s ∈ l'`, `Ioc (u n) (v n)` is eventually included + in `s`. + +This typeclass has the following “real” instances: `(a, pure a, ⊥)`, `(a, 𝓝[≥] a, 𝓝[>] a)`, +`(a, 𝓝[≤] a, 𝓝[≤] a)`, `(a, 𝓝 a, 𝓝 a)`. +Furthermore, we have the following instances that are equal to the previously mentioned instances: +`(a, 𝓝[{a}] a, ⊥)` and `(a, 𝓝[univ] a, 𝓝[univ] a)`. +While the difference between `Ici a` and `Ioi a` doesn't matter for theorems about Lebesgue measure, +it becomes important in the versions of FTC about any locally finite measure if this measure has an +atom at one of the endpoints. + +### Combining one-sided and two-sided derivatives + +There are some `FTC_filter` instances where the fact that it is one-sided or +two-sided depends on the point, namely `(x, 𝓝[Icc a b] x, 𝓝[Icc a b] x)` +(resp. `(x, 𝓝[[a, b]] x, 𝓝[[a, b]] x)`, where `[a, b] = set.uIcc a b`), +with `x ∈ Icc a b` (resp. `x ∈ [a, b]`). +This results in a two-sided derivatives for `x ∈ Ioo a b` and one-sided derivatives for +`x ∈ {a, b}`. Other instances could be added when needed (in that case, one also needs to add +instances for `filter.is_measurably_generated` and `filter.tendsto_Ixx_class`). + +## Tags + +integral, fundamental theorem of calculus, FTC-1, FTC-2, change of variables in integrals +-/ + +noncomputable theory +open topological_space (second_countable_topology) +open measure_theory set classical filter function + +open_locale classical topology filter ennreal big_operators interval nnreal + +variables {ι 𝕜 E F A : Type*} [normed_add_comm_group E] + [complete_space E] [normed_space ℝ E] + +namespace interval_integral + +/-! +### Fundamental theorem of calculus, part 1, for any measure + +In this section we prove a few lemmas that can be seen as versions of FTC-1 for interval integrals +w.r.t. any measure. Many theorems are formulated for one or two pairs of filters related by +`FTC_filter a l l'`. This typeclass has exactly four “real” instances: `(a, pure a, ⊥)`, +`(a, 𝓝[≥] a, 𝓝[>] a)`, `(a, 𝓝[≤] a, 𝓝[≤] a)`, `(a, 𝓝 a, 𝓝 a)`, and two instances +that are equal to the first and last “real” instances: `(a, 𝓝[{a}] a, ⊥)` and +`(a, 𝓝[univ] a, 𝓝[univ] a)`. We use this approach to avoid repeating arguments in many very similar +cases. Lean can automatically find both `a` and `l'` based on `l`. + +The most general theorem `measure_integral_sub_integral_sub_linear_is_o_of_tendsto_ae` can be seen +as a generalization of lemma `integral_has_strict_fderiv_at` below which states strict +differentiability of `∫ x in u..v, f x` in `(u, v)` at `(a, b)` for a measurable function `f` that +is integrable on `a..b` and is continuous at `a` and `b`. The lemma is generalized in three +directions: first, `measure_integral_sub_integral_sub_linear_is_o_of_tendsto_ae` deals with any +locally finite measure `μ`; second, it works for one-sided limits/derivatives; third, it assumes +only that `f` has finite limits almost surely at `a` and `b`. + +Namely, let `f` be a measurable function integrable on `a..b`. Let `(la, la')` be a pair of +`FTC_filter`s around `a`; let `(lb, lb')` be a pair of `FTC_filter`s around `b`. Suppose that `f` +has finite limits `ca` and `cb` at `la' ⊓ μ.ae` and `lb' ⊓ μ.ae`, respectively. Then +`∫ x in va..vb, f x ∂μ - ∫ x in ua..ub, f x ∂μ = ∫ x in ub..vb, cb ∂μ - ∫ x in ua..va, ca ∂μ + + o(‖∫ x in ua..va, (1:ℝ) ∂μ‖ + ‖∫ x in ub..vb, (1:ℝ) ∂μ‖)` +as `ua` and `va` tend to `la` while `ub` and `vb` tend to `lb`. + +This theorem is formulated with integral of constants instead of measures in the right hand sides +for two reasons: first, this way we avoid `min`/`max` in the statements; second, often it is +possible to write better `simp` lemmas for these integrals, see `integral_const` and +`integral_const_of_cdf`. + +In the next subsection we apply this theorem to prove various theorems about differentiability +of the integral w.r.t. Lebesgue measure. -/ + +/-- An auxiliary typeclass for the Fundamental theorem of calculus, part 1. It is used to formulate +theorems that work simultaneously for left and right one-sided derivatives of `∫ x in u..v, f x`. -/ +class FTC_filter (a : out_param ℝ) (outer : filter ℝ) (inner : out_param $ filter ℝ) + extends tendsto_Ixx_class Ioc outer inner : Prop := +(pure_le : pure a ≤ outer) +(le_nhds : inner ≤ 𝓝 a) +[meas_gen : is_measurably_generated inner] + +/- The `dangerous_instance` linter doesn't take `out_param`s into account, so it thinks that +`FTC_filter.to_tendsto_Ixx_class` is dangerous. Disable this linter using `nolint`. +-/ +attribute [nolint dangerous_instance] FTC_filter.to_tendsto_Ixx_class + +namespace FTC_filter + +instance pure (a : ℝ) : FTC_filter a (pure a) ⊥ := +{ pure_le := le_rfl, + le_nhds := bot_le } + +instance nhds_within_singleton (a : ℝ) : FTC_filter a (𝓝[{a}] a) ⊥ := +by { rw [nhds_within, principal_singleton, inf_eq_right.2 (pure_le_nhds a)], apply_instance } + +lemma finite_at_inner {a : ℝ} (l : filter ℝ) {l'} [h : FTC_filter a l l'] + {μ : measure ℝ} [is_locally_finite_measure μ] : + μ.finite_at_filter l' := +(μ.finite_at_nhds a).filter_mono h.le_nhds + +instance nhds (a : ℝ) : FTC_filter a (𝓝 a) (𝓝 a) := +{ pure_le := pure_le_nhds a, + le_nhds := le_rfl } + +instance nhds_univ (a : ℝ) : FTC_filter a (𝓝[univ] a) (𝓝 a) := +by { rw nhds_within_univ, apply_instance } + +instance nhds_left (a : ℝ) : FTC_filter a (𝓝[≤] a) (𝓝[≤] a) := +{ pure_le := pure_le_nhds_within right_mem_Iic, + le_nhds := inf_le_left } + +instance nhds_right (a : ℝ) : FTC_filter a (𝓝[≥] a) (𝓝[>] a) := +{ pure_le := pure_le_nhds_within left_mem_Ici, + le_nhds := inf_le_left } + +instance nhds_Icc {x a b : ℝ} [h : fact (x ∈ Icc a b)] : + FTC_filter x (𝓝[Icc a b] x) (𝓝[Icc a b] x) := +{ pure_le := pure_le_nhds_within h.out, + le_nhds := inf_le_left } + +instance nhds_uIcc {x a b : ℝ} [h : fact (x ∈ [a, b])] : + FTC_filter x (𝓝[[a, b]] x) (𝓝[[a, b]] x) := +by { haveI : fact (x ∈ set.Icc (min a b) (max a b)) := h, exact FTC_filter.nhds_Icc } + +end FTC_filter + +open asymptotics + +section + +variables {f : ℝ → E} {a b : ℝ} {c ca cb : E} {l l' la la' lb lb' : filter ℝ} {lt : filter ι} + {μ : measure ℝ} {u v ua va ub vb : ι → ℝ} + +/-- Fundamental theorem of calculus-1, local version for any measure. +Let filters `l` and `l'` be related by `tendsto_Ixx_class Ioc`. +If `f` has a finite limit `c` at `l' ⊓ μ.ae`, where `μ` is a measure +finite at `l'`, then `∫ x in u..v, f x ∂μ = ∫ x in u..v, c ∂μ + o(∫ x in u..v, 1 ∂μ)` as both +`u` and `v` tend to `l`. + +See also `measure_integral_sub_linear_is_o_of_tendsto_ae` for a version assuming +`[FTC_filter a l l']` and `[is_locally_finite_measure μ]`. If `l` is one of `𝓝[≥] a`, +`𝓝[≤] a`, `𝓝 a`, then it's easier to apply the non-primed version. +The primed version also works, e.g., for `l = l' = at_top`. + +We use integrals of constants instead of measures because this way it is easier to formulate +a statement that works in both cases `u ≤ v` and `v ≤ u`. -/ +lemma measure_integral_sub_linear_is_o_of_tendsto_ae' + [is_measurably_generated l'] [tendsto_Ixx_class Ioc l l'] + (hfm : strongly_measurable_at_filter f l' μ) (hf : tendsto f (l' ⊓ μ.ae) (𝓝 c)) + (hl : μ.finite_at_filter l') + (hu : tendsto u lt l) (hv : tendsto v lt l) : + (λ t, ∫ x in u t..v t, f x ∂μ - ∫ x in u t..v t, c ∂μ) =o[lt] (λ t, ∫ x in u t..v t, (1:ℝ) ∂μ) := +begin + have A := hf.integral_sub_linear_is_o_ae hfm hl (hu.Ioc hv), + have B := hf.integral_sub_linear_is_o_ae hfm hl (hv.Ioc hu), + simp only [integral_const'], + convert (A.trans_le _).sub (B.trans_le _), + { ext t, + simp_rw [interval_integral, sub_smul], + abel }, + all_goals { intro t, cases le_total (u t) (v t) with huv huv; simp [huv] } +end + +/-- Fundamental theorem of calculus-1, local version for any measure. +Let filters `l` and `l'` be related by `tendsto_Ixx_class Ioc`. +If `f` has a finite limit `c` at `l ⊓ μ.ae`, where `μ` is a measure +finite at `l`, then `∫ x in u..v, f x ∂μ = μ (Ioc u v) • c + o(μ(Ioc u v))` as both +`u` and `v` tend to `l` so that `u ≤ v`. + +See also `measure_integral_sub_linear_is_o_of_tendsto_ae_of_le` for a version assuming +`[FTC_filter a l l']` and `[is_locally_finite_measure μ]`. If `l` is one of `𝓝[≥] a`, +`𝓝[≤] a`, `𝓝 a`, then it's easier to apply the non-primed version. +The primed version also works, e.g., for `l = l' = at_top`. -/ +lemma measure_integral_sub_linear_is_o_of_tendsto_ae_of_le' + [is_measurably_generated l'] [tendsto_Ixx_class Ioc l l'] + (hfm : strongly_measurable_at_filter f l' μ) (hf : tendsto f (l' ⊓ μ.ae) (𝓝 c)) + (hl : μ.finite_at_filter l') (hu : tendsto u lt l) (hv : tendsto v lt l) (huv : u ≤ᶠ[lt] v) : + (λ t, ∫ x in u t..v t, f x ∂μ - (μ (Ioc (u t) (v t))).to_real • c) =o[lt] + (λ t, (μ $ Ioc (u t) (v t)).to_real) := +(measure_integral_sub_linear_is_o_of_tendsto_ae' hfm hf hl hu hv).congr' + (huv.mono $ λ x hx, by simp [integral_const', hx]) + (huv.mono $ λ x hx, by simp [integral_const', hx]) + +/-- Fundamental theorem of calculus-1, local version for any measure. +Let filters `l` and `l'` be related by `tendsto_Ixx_class Ioc`. +If `f` has a finite limit `c` at `l ⊓ μ.ae`, where `μ` is a measure +finite at `l`, then `∫ x in u..v, f x ∂μ = -μ (Ioc v u) • c + o(μ(Ioc v u))` as both +`u` and `v` tend to `l` so that `v ≤ u`. + +See also `measure_integral_sub_linear_is_o_of_tendsto_ae_of_ge` for a version assuming +`[FTC_filter a l l']` and `[is_locally_finite_measure μ]`. If `l` is one of `𝓝[≥] a`, +`𝓝[≤] a`, `𝓝 a`, then it's easier to apply the non-primed version. +The primed version also works, e.g., for `l = l' = at_top`. -/ +lemma measure_integral_sub_linear_is_o_of_tendsto_ae_of_ge' + [is_measurably_generated l'] [tendsto_Ixx_class Ioc l l'] + (hfm : strongly_measurable_at_filter f l' μ) (hf : tendsto f (l' ⊓ μ.ae) (𝓝 c)) + (hl : μ.finite_at_filter l') (hu : tendsto u lt l) (hv : tendsto v lt l) (huv : v ≤ᶠ[lt] u) : + (λ t, ∫ x in u t..v t, f x ∂μ + (μ (Ioc (v t) (u t))).to_real • c) =o[lt] + (λ t, (μ $ Ioc (v t) (u t)).to_real) := +(measure_integral_sub_linear_is_o_of_tendsto_ae_of_le' hfm hf hl hv hu huv).neg_left.congr_left $ + λ t, by simp [integral_symm (u t), add_comm] + +section + +variables [is_locally_finite_measure μ] [FTC_filter a l l'] + +include a + +local attribute [instance] FTC_filter.meas_gen + +/-- Fundamental theorem of calculus-1, local version for any measure. +Let filters `l` and `l'` be related by `[FTC_filter a l l']`; let `μ` be a locally finite measure. +If `f` has a finite limit `c` at `l' ⊓ μ.ae`, then +`∫ x in u..v, f x ∂μ = ∫ x in u..v, c ∂μ + o(∫ x in u..v, 1 ∂μ)` as both `u` and `v` tend to `l`. + +See also `measure_integral_sub_linear_is_o_of_tendsto_ae'` for a version that also works, e.g., for +`l = l' = at_top`. + +We use integrals of constants instead of measures because this way it is easier to formulate +a statement that works in both cases `u ≤ v` and `v ≤ u`. -/ +lemma measure_integral_sub_linear_is_o_of_tendsto_ae (hfm : strongly_measurable_at_filter f l' μ) + (hf : tendsto f (l' ⊓ μ.ae) (𝓝 c)) (hu : tendsto u lt l) (hv : tendsto v lt l) : + (λ t, ∫ x in u t..v t, f x ∂μ - ∫ x in u t..v t, c ∂μ) =o[lt] (λ t, ∫ x in u t..v t, (1:ℝ) ∂μ) := +measure_integral_sub_linear_is_o_of_tendsto_ae' hfm hf (FTC_filter.finite_at_inner l) hu hv + +/-- Fundamental theorem of calculus-1, local version for any measure. +Let filters `l` and `l'` be related by `[FTC_filter a l l']`; let `μ` be a locally finite measure. +If `f` has a finite limit `c` at `l' ⊓ μ.ae`, then +`∫ x in u..v, f x ∂μ = μ (Ioc u v) • c + o(μ(Ioc u v))` as both `u` and `v` tend to `l`. + +See also `measure_integral_sub_linear_is_o_of_tendsto_ae_of_le'` for a version that also works, +e.g., for `l = l' = at_top`. -/ +lemma measure_integral_sub_linear_is_o_of_tendsto_ae_of_le + (hfm : strongly_measurable_at_filter f l' μ) (hf : tendsto f (l' ⊓ μ.ae) (𝓝 c)) + (hu : tendsto u lt l) (hv : tendsto v lt l) (huv : u ≤ᶠ[lt] v) : + (λ t, ∫ x in u t..v t, f x ∂μ - (μ (Ioc (u t) (v t))).to_real • c) =o[lt] + (λ t, (μ $ Ioc (u t) (v t)).to_real) := +measure_integral_sub_linear_is_o_of_tendsto_ae_of_le' hfm hf (FTC_filter.finite_at_inner l) + hu hv huv + +/-- Fundamental theorem of calculus-1, local version for any measure. +Let filters `l` and `l'` be related by `[FTC_filter a l l']`; let `μ` be a locally finite measure. +If `f` has a finite limit `c` at `l' ⊓ μ.ae`, then +`∫ x in u..v, f x ∂μ = -μ (Ioc v u) • c + o(μ(Ioc v u))` as both `u` and `v` tend to `l`. + +See also `measure_integral_sub_linear_is_o_of_tendsto_ae_of_ge'` for a version that also works, +e.g., for `l = l' = at_top`. -/ +lemma measure_integral_sub_linear_is_o_of_tendsto_ae_of_ge + (hfm : strongly_measurable_at_filter f l' μ) (hf : tendsto f (l' ⊓ μ.ae) (𝓝 c)) + (hu : tendsto u lt l) (hv : tendsto v lt l) (huv : v ≤ᶠ[lt] u) : + (λ t, ∫ x in u t..v t, f x ∂μ + (μ (Ioc (v t) (u t))).to_real • c) =o[lt] + (λ t, (μ $ Ioc (v t) (u t)).to_real) := +measure_integral_sub_linear_is_o_of_tendsto_ae_of_ge' hfm hf (FTC_filter.finite_at_inner l) + hu hv huv + +end + +local attribute [instance] FTC_filter.meas_gen + +variables [FTC_filter a la la'] [FTC_filter b lb lb'] [is_locally_finite_measure μ] + +/-- Fundamental theorem of calculus-1, strict derivative in both limits for a locally finite +measure. + +Let `f` be a measurable function integrable on `a..b`. Let `(la, la')` be a pair of `FTC_filter`s +around `a`; let `(lb, lb')` be a pair of `FTC_filter`s around `b`. Suppose that `f` has finite +limits `ca` and `cb` at `la' ⊓ μ.ae` and `lb' ⊓ μ.ae`, respectively. +Then `∫ x in va..vb, f x ∂μ - ∫ x in ua..ub, f x ∂μ = + ∫ x in ub..vb, cb ∂μ - ∫ x in ua..va, ca ∂μ + + o(‖∫ x in ua..va, (1:ℝ) ∂μ‖ + ‖∫ x in ub..vb, (1:ℝ) ∂μ‖)` +as `ua` and `va` tend to `la` while `ub` and `vb` tend to `lb`. +-/ +lemma measure_integral_sub_integral_sub_linear_is_o_of_tendsto_ae + (hab : interval_integrable f μ a b) + (hmeas_a : strongly_measurable_at_filter f la' μ) + (hmeas_b : strongly_measurable_at_filter f lb' μ) + (ha_lim : tendsto f (la' ⊓ μ.ae) (𝓝 ca)) (hb_lim : tendsto f (lb' ⊓ μ.ae) (𝓝 cb)) + (hua : tendsto ua lt la) (hva : tendsto va lt la) + (hub : tendsto ub lt lb) (hvb : tendsto vb lt lb) : + (λ t, (∫ x in va t..vb t, f x ∂μ) - (∫ x in ua t..ub t, f x ∂μ) - + (∫ x in ub t..vb t, cb ∂μ - ∫ x in ua t..va t, ca ∂μ)) =o[lt] + (λ t, ‖∫ x in ua t..va t, (1:ℝ) ∂μ‖ + ‖∫ x in ub t..vb t, (1:ℝ) ∂μ‖) := +begin + refine + ((measure_integral_sub_linear_is_o_of_tendsto_ae hmeas_a ha_lim hua hva).neg_left.add_add + (measure_integral_sub_linear_is_o_of_tendsto_ae hmeas_b hb_lim hub hvb)).congr' + _ eventually_eq.rfl, + have A : ∀ᶠ t in lt, interval_integrable f μ (ua t) (va t) := + ha_lim.eventually_interval_integrable_ae hmeas_a (FTC_filter.finite_at_inner la) hua hva, + have A' : ∀ᶠ t in lt, interval_integrable f μ a (ua t) := + ha_lim.eventually_interval_integrable_ae hmeas_a (FTC_filter.finite_at_inner la) + (tendsto_const_pure.mono_right FTC_filter.pure_le) hua, + have B : ∀ᶠ t in lt, interval_integrable f μ (ub t) (vb t) := + hb_lim.eventually_interval_integrable_ae hmeas_b (FTC_filter.finite_at_inner lb) hub hvb, + have B' : ∀ᶠ t in lt, interval_integrable f μ b (ub t) := + hb_lim.eventually_interval_integrable_ae hmeas_b (FTC_filter.finite_at_inner lb) + (tendsto_const_pure.mono_right FTC_filter.pure_le) hub, + filter_upwards [A, A', B, B'] with _ ua_va a_ua ub_vb b_ub, + rw [← integral_interval_sub_interval_comm'], + { dsimp only [], abel, }, + exacts [ub_vb, ua_va, b_ub.symm.trans $ hab.symm.trans a_ua] +end + +/-- Fundamental theorem of calculus-1, strict derivative in right endpoint for a locally finite +measure. + +Let `f` be a measurable function integrable on `a..b`. Let `(lb, lb')` be a pair of `FTC_filter`s +around `b`. Suppose that `f` has a finite limit `c` at `lb' ⊓ μ.ae`. + +Then `∫ x in a..v, f x ∂μ - ∫ x in a..u, f x ∂μ = ∫ x in u..v, c ∂μ + o(∫ x in u..v, (1:ℝ) ∂μ)` +as `u` and `v` tend to `lb`. +-/ +lemma measure_integral_sub_integral_sub_linear_is_o_of_tendsto_ae_right + (hab : interval_integrable f μ a b) (hmeas : strongly_measurable_at_filter f lb' μ) + (hf : tendsto f (lb' ⊓ μ.ae) (𝓝 c)) (hu : tendsto u lt lb) (hv : tendsto v lt lb) : + (λ t, ∫ x in a..v t, f x ∂μ - ∫ x in a..u t, f x ∂μ - ∫ x in u t..v t, c ∂μ) =o[lt] + (λ t, ∫ x in u t..v t, (1:ℝ) ∂μ) := +by simpa using measure_integral_sub_integral_sub_linear_is_o_of_tendsto_ae + hab strongly_measurable_at_bot hmeas ((tendsto_bot : tendsto _ ⊥ (𝓝 0)).mono_left inf_le_left) + hf (tendsto_const_pure : tendsto _ _ (pure a)) tendsto_const_pure hu hv + +/-- Fundamental theorem of calculus-1, strict derivative in left endpoint for a locally finite +measure. + +Let `f` be a measurable function integrable on `a..b`. Let `(la, la')` be a pair of `FTC_filter`s +around `a`. Suppose that `f` has a finite limit `c` at `la' ⊓ μ.ae`. + +Then `∫ x in v..b, f x ∂μ - ∫ x in u..b, f x ∂μ = -∫ x in u..v, c ∂μ + o(∫ x in u..v, (1:ℝ) ∂μ)` +as `u` and `v` tend to `la`. +-/ +lemma measure_integral_sub_integral_sub_linear_is_o_of_tendsto_ae_left + (hab : interval_integrable f μ a b) (hmeas : strongly_measurable_at_filter f la' μ) + (hf : tendsto f (la' ⊓ μ.ae) (𝓝 c)) (hu : tendsto u lt la) (hv : tendsto v lt la) : + (λ t, ∫ x in v t..b, f x ∂μ - ∫ x in u t..b, f x ∂μ + ∫ x in u t..v t, c ∂μ) =o[lt] + (λ t, ∫ x in u t..v t, (1:ℝ) ∂μ) := +by simpa using measure_integral_sub_integral_sub_linear_is_o_of_tendsto_ae + hab hmeas strongly_measurable_at_bot hf ((tendsto_bot : tendsto _ ⊥ (𝓝 0)).mono_left inf_le_left) + hu hv (tendsto_const_pure : tendsto _ _ (pure b)) tendsto_const_pure + +end + +/-! +### Fundamental theorem of calculus-1 for Lebesgue measure + +In this section we restate theorems from the previous section for Lebesgue measure. +In particular, we prove that `∫ x in u..v, f x` is strictly differentiable in `(u, v)` +at `(a, b)` provided that `f` is integrable on `a..b` and is continuous at `a` and `b`. +-/ + +variables {f : ℝ → E} {c ca cb : E} {l l' la la' lb lb' : filter ℝ} {lt : filter ι} + {a b z : ℝ} {u v ua ub va vb : ι → ℝ} [FTC_filter a la la'] [FTC_filter b lb lb'] + +/-! +#### Auxiliary `is_o` statements + +In this section we prove several lemmas that can be interpreted as strict differentiability of +`(u, v) ↦ ∫ x in u..v, f x ∂μ` in `u` and/or `v` at a filter. The statements use `is_o` because +we have no definition of `has_strict_(f)deriv_at_filter` in the library. +-/ + +/-- Fundamental theorem of calculus-1, local version. If `f` has a finite limit `c` almost surely at +`l'`, where `(l, l')` is an `FTC_filter` pair around `a`, then +`∫ x in u..v, f x ∂μ = (v - u) • c + o (v - u)` as both `u` and `v` tend to `l`. -/ +lemma integral_sub_linear_is_o_of_tendsto_ae [FTC_filter a l l'] + (hfm : strongly_measurable_at_filter f l') (hf : tendsto f (l' ⊓ volume.ae) (𝓝 c)) + {u v : ι → ℝ} (hu : tendsto u lt l) (hv : tendsto v lt l) : + (λ t, (∫ x in u t..v t, f x) - (v t - u t) • c) =o[lt] (v - u) := +by simpa [integral_const] using measure_integral_sub_linear_is_o_of_tendsto_ae hfm hf hu hv + +/-- Fundamental theorem of calculus-1, strict differentiability at filter in both endpoints. +If `f` is a measurable function integrable on `a..b`, `(la, la')` is an `FTC_filter` pair around +`a`, and `(lb, lb')` is an `FTC_filter` pair around `b`, and `f` has finite limits `ca` and `cb` +almost surely at `la'` and `lb'`, respectively, then +`(∫ x in va..vb, f x) - ∫ x in ua..ub, f x = (vb - ub) • cb - (va - ua) • ca + + o(‖va - ua‖ + ‖vb - ub‖)` as `ua` and `va` tend to `la` while `ub` and `vb` tend to `lb`. + +This lemma could've been formulated using `has_strict_fderiv_at_filter` if we had this +definition. -/ +lemma integral_sub_integral_sub_linear_is_o_of_tendsto_ae + (hab : interval_integrable f volume a b) + (hmeas_a : strongly_measurable_at_filter f la') (hmeas_b : strongly_measurable_at_filter f lb') + (ha_lim : tendsto f (la' ⊓ volume.ae) (𝓝 ca)) (hb_lim : tendsto f (lb' ⊓ volume.ae) (𝓝 cb)) + (hua : tendsto ua lt la) (hva : tendsto va lt la) + (hub : tendsto ub lt lb) (hvb : tendsto vb lt lb) : + (λ t, (∫ x in va t..vb t, f x) - (∫ x in ua t..ub t, f x) - + ((vb t - ub t) • cb - (va t - ua t) • ca)) =o[lt] (λ t, ‖va t - ua t‖ + ‖vb t - ub t‖) := +by simpa [integral_const] + using measure_integral_sub_integral_sub_linear_is_o_of_tendsto_ae hab hmeas_a hmeas_b + ha_lim hb_lim hua hva hub hvb + +/-- Fundamental theorem of calculus-1, strict differentiability at filter in both endpoints. +If `f` is a measurable function integrable on `a..b`, `(lb, lb')` is an `FTC_filter` pair +around `b`, and `f` has a finite limit `c` almost surely at `lb'`, then +`(∫ x in a..v, f x) - ∫ x in a..u, f x = (v - u) • c + o(‖v - u‖)` as `u` and `v` tend to `lb`. + +This lemma could've been formulated using `has_strict_deriv_at_filter` if we had this definition. -/ +lemma integral_sub_integral_sub_linear_is_o_of_tendsto_ae_right + (hab : interval_integrable f volume a b) (hmeas : strongly_measurable_at_filter f lb') + (hf : tendsto f (lb' ⊓ volume.ae) (𝓝 c)) (hu : tendsto u lt lb) (hv : tendsto v lt lb) : + (λ t, (∫ x in a..v t, f x) - (∫ x in a..u t, f x) - (v t - u t) • c) =o[lt] (v - u) := +by simpa only [integral_const, smul_eq_mul, mul_one] using + measure_integral_sub_integral_sub_linear_is_o_of_tendsto_ae_right hab hmeas hf hu hv + +/-- Fundamental theorem of calculus-1, strict differentiability at filter in both endpoints. +If `f` is a measurable function integrable on `a..b`, `(la, la')` is an `FTC_filter` pair +around `a`, and `f` has a finite limit `c` almost surely at `la'`, then +`(∫ x in v..b, f x) - ∫ x in u..b, f x = -(v - u) • c + o(‖v - u‖)` as `u` and `v` tend to `la`. + +This lemma could've been formulated using `has_strict_deriv_at_filter` if we had this definition. -/ +lemma integral_sub_integral_sub_linear_is_o_of_tendsto_ae_left + (hab : interval_integrable f volume a b) (hmeas : strongly_measurable_at_filter f la') + (hf : tendsto f (la' ⊓ volume.ae) (𝓝 c)) (hu : tendsto u lt la) (hv : tendsto v lt la) : + (λ t, (∫ x in v t..b, f x) - (∫ x in u t..b, f x) + (v t - u t) • c) =o[lt] (v - u) := +by simpa only [integral_const, smul_eq_mul, mul_one] using + measure_integral_sub_integral_sub_linear_is_o_of_tendsto_ae_left hab hmeas hf hu hv + +open continuous_linear_map (fst snd smul_right sub_apply smul_right_apply coe_fst' coe_snd' map_sub) + +/-! +#### Strict differentiability + +In this section we prove that for a measurable function `f` integrable on `a..b`, + +* `integral_has_strict_fderiv_at_of_tendsto_ae`: the function `(u, v) ↦ ∫ x in u..v, f x` has + derivative `(u, v) ↦ v • cb - u • ca` at `(a, b)` in the sense of strict differentiability + provided that `f` tends to `ca` and `cb` almost surely as `x` tendsto to `a` and `b`, + respectively; + +* `integral_has_strict_fderiv_at`: the function `(u, v) ↦ ∫ x in u..v, f x` has + derivative `(u, v) ↦ v • f b - u • f a` at `(a, b)` in the sense of strict differentiability + provided that `f` is continuous at `a` and `b`; + +* `integral_has_strict_deriv_at_of_tendsto_ae_right`: the function `u ↦ ∫ x in a..u, f x` has + derivative `c` at `b` in the sense of strict differentiability provided that `f` tends to `c` + almost surely as `x` tends to `b`; + +* `integral_has_strict_deriv_at_right`: the function `u ↦ ∫ x in a..u, f x` has derivative `f b` at + `b` in the sense of strict differentiability provided that `f` is continuous at `b`; + +* `integral_has_strict_deriv_at_of_tendsto_ae_left`: the function `u ↦ ∫ x in u..b, f x` has + derivative `-c` at `a` in the sense of strict differentiability provided that `f` tends to `c` + almost surely as `x` tends to `a`; + +* `integral_has_strict_deriv_at_left`: the function `u ↦ ∫ x in u..b, f x` has derivative `-f a` at + `a` in the sense of strict differentiability provided that `f` is continuous at `a`. +-/ + +/-- Fundamental theorem of calculus-1: if `f : ℝ → E` is integrable on `a..b` and `f x` has finite +limits `ca` and `cb` almost surely as `x` tends to `a` and `b`, respectively, then +`(u, v) ↦ ∫ x in u..v, f x` has derivative `(u, v) ↦ v • cb - u • ca` at `(a, b)` +in the sense of strict differentiability. -/ +lemma integral_has_strict_fderiv_at_of_tendsto_ae + (hf : interval_integrable f volume a b) + (hmeas_a : strongly_measurable_at_filter f (𝓝 a)) + (hmeas_b : strongly_measurable_at_filter f (𝓝 b)) + (ha : tendsto f (𝓝 a ⊓ volume.ae) (𝓝 ca)) (hb : tendsto f (𝓝 b ⊓ volume.ae) (𝓝 cb)) : + has_strict_fderiv_at (λ p : ℝ × ℝ, ∫ x in p.1..p.2, f x) + ((snd ℝ ℝ ℝ).smul_right cb - (fst ℝ ℝ ℝ).smul_right ca) (a, b) := +begin + have := integral_sub_integral_sub_linear_is_o_of_tendsto_ae hf hmeas_a hmeas_b ha hb + ((continuous_fst.comp continuous_snd).tendsto ((a, b), (a, b))) + ((continuous_fst.comp continuous_fst).tendsto ((a, b), (a, b))) + ((continuous_snd.comp continuous_snd).tendsto ((a, b), (a, b))) + ((continuous_snd.comp continuous_fst).tendsto ((a, b), (a, b))), + refine (this.congr_left _).trans_is_O _, + { intro x, simp [sub_smul] }, + { exact is_O_fst_prod.norm_left.add is_O_snd_prod.norm_left } +end + +/-- Fundamental theorem of calculus-1: if `f : ℝ → E` is integrable on `a..b` and `f` is continuous +at `a` and `b`, then `(u, v) ↦ ∫ x in u..v, f x` has derivative `(u, v) ↦ v • cb - u • ca` +at `(a, b)` in the sense of strict differentiability. -/ +lemma integral_has_strict_fderiv_at + (hf : interval_integrable f volume a b) + (hmeas_a : strongly_measurable_at_filter f (𝓝 a)) + (hmeas_b : strongly_measurable_at_filter f (𝓝 b)) + (ha : continuous_at f a) (hb : continuous_at f b) : + has_strict_fderiv_at (λ p : ℝ × ℝ, ∫ x in p.1..p.2, f x) + ((snd ℝ ℝ ℝ).smul_right (f b) - (fst ℝ ℝ ℝ).smul_right (f a)) (a, b) := +integral_has_strict_fderiv_at_of_tendsto_ae hf hmeas_a hmeas_b + (ha.mono_left inf_le_left) (hb.mono_left inf_le_left) + +/-- **First Fundamental Theorem of Calculus**: if `f : ℝ → E` is integrable on `a..b` and `f x` has +a finite limit `c` almost surely at `b`, then `u ↦ ∫ x in a..u, f x` has derivative `c` at `b` in +the sense of strict differentiability. -/ +lemma integral_has_strict_deriv_at_of_tendsto_ae_right + (hf : interval_integrable f volume a b) (hmeas : strongly_measurable_at_filter f (𝓝 b)) + (hb : tendsto f (𝓝 b ⊓ volume.ae) (𝓝 c)) : has_strict_deriv_at (λ u, ∫ x in a..u, f x) c b := +integral_sub_integral_sub_linear_is_o_of_tendsto_ae_right hf hmeas hb continuous_at_snd + continuous_at_fst + +/-- Fundamental theorem of calculus-1: if `f : ℝ → E` is integrable on `a..b` and `f` is continuous +at `b`, then `u ↦ ∫ x in a..u, f x` has derivative `f b` at `b` in the sense of strict +differentiability. -/ +lemma integral_has_strict_deriv_at_right + (hf : interval_integrable f volume a b) (hmeas : strongly_measurable_at_filter f (𝓝 b)) + (hb : continuous_at f b) : has_strict_deriv_at (λ u, ∫ x in a..u, f x) (f b) b := +integral_has_strict_deriv_at_of_tendsto_ae_right hf hmeas (hb.mono_left inf_le_left) + +/-- Fundamental theorem of calculus-1: if `f : ℝ → E` is integrable on `a..b` and `f x` has a finite +limit `c` almost surely at `a`, then `u ↦ ∫ x in u..b, f x` has derivative `-c` at `a` in the sense +of strict differentiability. -/ +lemma integral_has_strict_deriv_at_of_tendsto_ae_left + (hf : interval_integrable f volume a b) (hmeas : strongly_measurable_at_filter f (𝓝 a)) + (ha : tendsto f (𝓝 a ⊓ volume.ae) (𝓝 c)) : has_strict_deriv_at (λ u, ∫ x in u..b, f x) (-c) a := +by simpa only [← integral_symm] + using (integral_has_strict_deriv_at_of_tendsto_ae_right hf.symm hmeas ha).neg + +/-- Fundamental theorem of calculus-1: if `f : ℝ → E` is integrable on `a..b` and `f` is continuous +at `a`, then `u ↦ ∫ x in u..b, f x` has derivative `-f a` at `a` in the sense of strict +differentiability. -/ +lemma integral_has_strict_deriv_at_left + (hf : interval_integrable f volume a b) (hmeas : strongly_measurable_at_filter f (𝓝 a)) + (ha : continuous_at f a) : has_strict_deriv_at (λ u, ∫ x in u..b, f x) (-f a) a := +by simpa only [← integral_symm] using (integral_has_strict_deriv_at_right hf.symm hmeas ha).neg + +/-- Fundamental theorem of calculus-1: if `f : ℝ → E` is continuous, then `u ↦ ∫ x in a..u, f x` +has derivative `f b` at `b` in the sense of strict differentiability. -/ +lemma _root_.continuous.integral_has_strict_deriv_at {f : ℝ → E} (hf : continuous f) (a b : ℝ) : + has_strict_deriv_at (λ u, ∫ (x : ℝ) in a..u, f x) (f b) b := +integral_has_strict_deriv_at_right (hf.interval_integrable _ _) + (hf.strongly_measurable_at_filter _ _) hf.continuous_at + +/-- Fundamental theorem of calculus-1: if `f : ℝ → E` is continuous, then the derivative +of `u ↦ ∫ x in a..u, f x` at `b` is `f b`. -/ +lemma _root_.continuous.deriv_integral (f : ℝ → E) (hf : continuous f) (a b : ℝ) : + deriv (λ u, ∫ (x : ℝ) in a..u, f x) b = f b := +(hf.integral_has_strict_deriv_at a b).has_deriv_at.deriv + +/-! +#### Fréchet differentiability + +In this subsection we restate results from the previous subsection in terms of `has_fderiv_at`, +`has_deriv_at`, `fderiv`, and `deriv`. +-/ + +/-- Fundamental theorem of calculus-1: if `f : ℝ → E` is integrable on `a..b` and `f x` has finite +limits `ca` and `cb` almost surely as `x` tends to `a` and `b`, respectively, then +`(u, v) ↦ ∫ x in u..v, f x` has derivative `(u, v) ↦ v • cb - u • ca` at `(a, b)`. -/ +lemma integral_has_fderiv_at_of_tendsto_ae (hf : interval_integrable f volume a b) + (hmeas_a : strongly_measurable_at_filter f (𝓝 a)) + (hmeas_b : strongly_measurable_at_filter f (𝓝 b)) + (ha : tendsto f (𝓝 a ⊓ volume.ae) (𝓝 ca)) (hb : tendsto f (𝓝 b ⊓ volume.ae) (𝓝 cb)) : + has_fderiv_at (λ p : ℝ × ℝ, ∫ x in p.1..p.2, f x) + ((snd ℝ ℝ ℝ).smul_right cb - (fst ℝ ℝ ℝ).smul_right ca) (a, b) := +(integral_has_strict_fderiv_at_of_tendsto_ae hf hmeas_a hmeas_b ha hb).has_fderiv_at + +/-- Fundamental theorem of calculus-1: if `f : ℝ → E` is integrable on `a..b` and `f` is continuous +at `a` and `b`, then `(u, v) ↦ ∫ x in u..v, f x` has derivative `(u, v) ↦ v • cb - u • ca` +at `(a, b)`. -/ +lemma integral_has_fderiv_at (hf : interval_integrable f volume a b) + (hmeas_a : strongly_measurable_at_filter f (𝓝 a)) + (hmeas_b : strongly_measurable_at_filter f (𝓝 b)) + (ha : continuous_at f a) (hb : continuous_at f b) : + has_fderiv_at (λ p : ℝ × ℝ, ∫ x in p.1..p.2, f x) + ((snd ℝ ℝ ℝ).smul_right (f b) - (fst ℝ ℝ ℝ).smul_right (f a)) (a, b) := +(integral_has_strict_fderiv_at hf hmeas_a hmeas_b ha hb).has_fderiv_at + +/-- Fundamental theorem of calculus-1: if `f : ℝ → E` is integrable on `a..b` and `f x` has finite +limits `ca` and `cb` almost surely as `x` tends to `a` and `b`, respectively, then `fderiv` +derivative of `(u, v) ↦ ∫ x in u..v, f x` at `(a, b)` equals `(u, v) ↦ v • cb - u • ca`. -/ +lemma fderiv_integral_of_tendsto_ae (hf : interval_integrable f volume a b) + (hmeas_a : strongly_measurable_at_filter f (𝓝 a)) + (hmeas_b : strongly_measurable_at_filter f (𝓝 b)) + (ha : tendsto f (𝓝 a ⊓ volume.ae) (𝓝 ca)) (hb : tendsto f (𝓝 b ⊓ volume.ae) (𝓝 cb)) : + fderiv ℝ (λ p : ℝ × ℝ, ∫ x in p.1..p.2, f x) (a, b) = + (snd ℝ ℝ ℝ).smul_right cb - (fst ℝ ℝ ℝ).smul_right ca := +(integral_has_fderiv_at_of_tendsto_ae hf hmeas_a hmeas_b ha hb).fderiv + +/-- Fundamental theorem of calculus-1: if `f : ℝ → E` is integrable on `a..b` and `f` is continuous +at `a` and `b`, then `fderiv` derivative of `(u, v) ↦ ∫ x in u..v, f x` at `(a, b)` equals `(u, v) ↦ +v • cb - u • ca`. -/ +lemma fderiv_integral (hf : interval_integrable f volume a b) + (hmeas_a : strongly_measurable_at_filter f (𝓝 a)) + (hmeas_b : strongly_measurable_at_filter f (𝓝 b)) + (ha : continuous_at f a) (hb : continuous_at f b) : + fderiv ℝ (λ p : ℝ × ℝ, ∫ x in p.1..p.2, f x) (a, b) = + (snd ℝ ℝ ℝ).smul_right (f b) - (fst ℝ ℝ ℝ).smul_right (f a) := +(integral_has_fderiv_at hf hmeas_a hmeas_b ha hb).fderiv + +/-- Fundamental theorem of calculus-1: if `f : ℝ → E` is integrable on `a..b` and `f x` has a finite +limit `c` almost surely at `b`, then `u ↦ ∫ x in a..u, f x` has derivative `c` at `b`. -/ +lemma integral_has_deriv_at_of_tendsto_ae_right + (hf : interval_integrable f volume a b) (hmeas : strongly_measurable_at_filter f (𝓝 b)) + (hb : tendsto f (𝓝 b ⊓ volume.ae) (𝓝 c)) : has_deriv_at (λ u, ∫ x in a..u, f x) c b := +(integral_has_strict_deriv_at_of_tendsto_ae_right hf hmeas hb).has_deriv_at + +/-- Fundamental theorem of calculus-1: if `f : ℝ → E` is integrable on `a..b` and `f` is continuous +at `b`, then `u ↦ ∫ x in a..u, f x` has derivative `f b` at `b`. -/ +lemma integral_has_deriv_at_right + (hf : interval_integrable f volume a b) (hmeas : strongly_measurable_at_filter f (𝓝 b)) + (hb : continuous_at f b) : has_deriv_at (λ u, ∫ x in a..u, f x) (f b) b := +(integral_has_strict_deriv_at_right hf hmeas hb).has_deriv_at + +/-- Fundamental theorem of calculus: if `f : ℝ → E` is integrable on `a..b` and `f` has a finite +limit `c` almost surely at `b`, then the derivative of `u ↦ ∫ x in a..u, f x` at `b` equals `c`. -/ +lemma deriv_integral_of_tendsto_ae_right + (hf : interval_integrable f volume a b) (hmeas : strongly_measurable_at_filter f (𝓝 b)) + (hb : tendsto f (𝓝 b ⊓ volume.ae) (𝓝 c)) : deriv (λ u, ∫ x in a..u, f x) b = c := +(integral_has_deriv_at_of_tendsto_ae_right hf hmeas hb).deriv + +/-- Fundamental theorem of calculus: if `f : ℝ → E` is integrable on `a..b` and `f` is continuous +at `b`, then the derivative of `u ↦ ∫ x in a..u, f x` at `b` equals `f b`. -/ +lemma deriv_integral_right + (hf : interval_integrable f volume a b) (hmeas : strongly_measurable_at_filter f (𝓝 b)) + (hb : continuous_at f b) : + deriv (λ u, ∫ x in a..u, f x) b = f b := +(integral_has_deriv_at_right hf hmeas hb).deriv + +/-- Fundamental theorem of calculus-1: if `f : ℝ → E` is integrable on `a..b` and `f x` has a finite +limit `c` almost surely at `a`, then `u ↦ ∫ x in u..b, f x` has derivative `-c` at `a`. -/ +lemma integral_has_deriv_at_of_tendsto_ae_left + (hf : interval_integrable f volume a b) (hmeas : strongly_measurable_at_filter f (𝓝 a)) + (ha : tendsto f (𝓝 a ⊓ volume.ae) (𝓝 c)) : has_deriv_at (λ u, ∫ x in u..b, f x) (-c) a := +(integral_has_strict_deriv_at_of_tendsto_ae_left hf hmeas ha).has_deriv_at + +/-- Fundamental theorem of calculus-1: if `f : ℝ → E` is integrable on `a..b` and `f` is continuous +at `a`, then `u ↦ ∫ x in u..b, f x` has derivative `-f a` at `a`. -/ +lemma integral_has_deriv_at_left + (hf : interval_integrable f volume a b) (hmeas : strongly_measurable_at_filter f (𝓝 a)) + (ha : continuous_at f a) : + has_deriv_at (λ u, ∫ x in u..b, f x) (-f a) a := +(integral_has_strict_deriv_at_left hf hmeas ha).has_deriv_at + +/-- Fundamental theorem of calculus: if `f : ℝ → E` is integrable on `a..b` and `f` has a finite +limit `c` almost surely at `a`, then the derivative of `u ↦ ∫ x in u..b, f x` at `a` equals `-c`. -/ +lemma deriv_integral_of_tendsto_ae_left + (hf : interval_integrable f volume a b) (hmeas : strongly_measurable_at_filter f (𝓝 a)) + (hb : tendsto f (𝓝 a ⊓ volume.ae) (𝓝 c)) : deriv (λ u, ∫ x in u..b, f x) a = -c := +(integral_has_deriv_at_of_tendsto_ae_left hf hmeas hb).deriv + +/-- Fundamental theorem of calculus: if `f : ℝ → E` is integrable on `a..b` and `f` is continuous +at `a`, then the derivative of `u ↦ ∫ x in u..b, f x` at `a` equals `-f a`. -/ +lemma deriv_integral_left + (hf : interval_integrable f volume a b) (hmeas : strongly_measurable_at_filter f (𝓝 a)) + (hb : continuous_at f a) : + deriv (λ u, ∫ x in u..b, f x) a = -f a := +(integral_has_deriv_at_left hf hmeas hb).deriv + +/-! +#### One-sided derivatives +-/ + +/-- Let `f` be a measurable function integrable on `a..b`. The function `(u, v) ↦ ∫ x in u..v, f x` +has derivative `(u, v) ↦ v • cb - u • ca` within `s × t` at `(a, b)`, where +`s ∈ {Iic a, {a}, Ici a, univ}` and `t ∈ {Iic b, {b}, Ici b, univ}` provided that `f` tends to `ca` +and `cb` almost surely at the filters `la` and `lb` from the following table. + +| `s` | `la` | `t` | `lb` | +| ------- | ---- | --- | ---- | +| `Iic a` | `𝓝[≤] a` | `Iic b` | `𝓝[≤] b` | +| `Ici a` | `𝓝[>] a` | `Ici b` | `𝓝[>] b` | +| `{a}` | `⊥` | `{b}` | `⊥` | +| `univ` | `𝓝 a` | `univ` | `𝓝 b` | +-/ +lemma integral_has_fderiv_within_at_of_tendsto_ae + (hf : interval_integrable f volume a b) + {s t : set ℝ} [FTC_filter a (𝓝[s] a) la] [FTC_filter b (𝓝[t] b) lb] + (hmeas_a : strongly_measurable_at_filter f la) (hmeas_b : strongly_measurable_at_filter f lb) + (ha : tendsto f (la ⊓ volume.ae) (𝓝 ca)) (hb : tendsto f (lb ⊓ volume.ae) (𝓝 cb)) : + has_fderiv_within_at (λ p : ℝ × ℝ, ∫ x in p.1..p.2, f x) + ((snd ℝ ℝ ℝ).smul_right cb - (fst ℝ ℝ ℝ).smul_right ca) (s ×ˢ t) (a, b) := +begin + rw [has_fderiv_within_at, nhds_within_prod_eq], + have := integral_sub_integral_sub_linear_is_o_of_tendsto_ae hf hmeas_a hmeas_b ha hb + (tendsto_const_pure.mono_right FTC_filter.pure_le : tendsto _ _ (𝓝[s] a)) tendsto_fst + (tendsto_const_pure.mono_right FTC_filter.pure_le : tendsto _ _ (𝓝[t] b)) tendsto_snd, + refine (this.congr_left _).trans_is_O _, + { intro x, simp [sub_smul] }, + { exact is_O_fst_prod.norm_left.add is_O_snd_prod.norm_left } +end + +/-- Let `f` be a measurable function integrable on `a..b`. The function `(u, v) ↦ ∫ x in u..v, f x` +has derivative `(u, v) ↦ v • f b - u • f a` within `s × t` at `(a, b)`, where +`s ∈ {Iic a, {a}, Ici a, univ}` and `t ∈ {Iic b, {b}, Ici b, univ}` provided that `f` tends to +`f a` and `f b` at the filters `la` and `lb` from the following table. In most cases this assumption +is definitionally equal `continuous_at f _` or `continuous_within_at f _ _`. + +| `s` | `la` | `t` | `lb` | +| ------- | ---- | --- | ---- | +| `Iic a` | `𝓝[≤] a` | `Iic b` | `𝓝[≤] b` | +| `Ici a` | `𝓝[>] a` | `Ici b` | `𝓝[>] b` | +| `{a}` | `⊥` | `{b}` | `⊥` | +| `univ` | `𝓝 a` | `univ` | `𝓝 b` | +-/ +lemma integral_has_fderiv_within_at + (hf : interval_integrable f volume a b) + (hmeas_a : strongly_measurable_at_filter f la) (hmeas_b : strongly_measurable_at_filter f lb) + {s t : set ℝ} [FTC_filter a (𝓝[s] a) la] [FTC_filter b (𝓝[t] b) lb] + (ha : tendsto f la (𝓝 $ f a)) (hb : tendsto f lb (𝓝 $ f b)) : + has_fderiv_within_at (λ p : ℝ × ℝ, ∫ x in p.1..p.2, f x) + ((snd ℝ ℝ ℝ).smul_right (f b) - (fst ℝ ℝ ℝ).smul_right (f a)) (s ×ˢ t) (a, b) := +integral_has_fderiv_within_at_of_tendsto_ae hf hmeas_a hmeas_b (ha.mono_left inf_le_left) + (hb.mono_left inf_le_left) + +/-- An auxiliary tactic closing goals `unique_diff_within_at ℝ s a` where +`s ∈ {Iic a, Ici a, univ}`. -/ +meta def unique_diff_within_at_Ici_Iic_univ : tactic unit := +`[apply_rules [unique_diff_on.unique_diff_within_at, unique_diff_on_Ici, unique_diff_on_Iic, + left_mem_Ici, right_mem_Iic, unique_diff_within_at_univ]] + +/-- Let `f` be a measurable function integrable on `a..b`. Choose `s ∈ {Iic a, Ici a, univ}` +and `t ∈ {Iic b, Ici b, univ}`. Suppose that `f` tends to `ca` and `cb` almost surely at the filters +`la` and `lb` from the table below. Then `fderiv_within ℝ (λ p, ∫ x in p.1..p.2, f x) (s ×ˢ t)` +is equal to `(u, v) ↦ u • cb - v • ca`. + +| `s` | `la` | `t` | `lb` | +| ------- | ---- | --- | ---- | +| `Iic a` | `𝓝[≤] a` | `Iic b` | `𝓝[≤] b` | +| `Ici a` | `𝓝[>] a` | `Ici b` | `𝓝[>] b` | +| `{a}` | `⊥` | `{b}` | `⊥` | +| `univ` | `𝓝 a` | `univ` | `𝓝 b` | +-/ +lemma fderiv_within_integral_of_tendsto_ae + (hf : interval_integrable f volume a b) + (hmeas_a : strongly_measurable_at_filter f la) (hmeas_b : strongly_measurable_at_filter f lb) + {s t : set ℝ} [FTC_filter a (𝓝[s] a) la] [FTC_filter b (𝓝[t] b) lb] + (ha : tendsto f (la ⊓ volume.ae) (𝓝 ca)) (hb : tendsto f (lb ⊓ volume.ae) (𝓝 cb)) + (hs : unique_diff_within_at ℝ s a . unique_diff_within_at_Ici_Iic_univ) + (ht : unique_diff_within_at ℝ t b . unique_diff_within_at_Ici_Iic_univ) : + fderiv_within ℝ (λ p : ℝ × ℝ, ∫ x in p.1..p.2, f x) (s ×ˢ t) (a, b) = + ((snd ℝ ℝ ℝ).smul_right cb - (fst ℝ ℝ ℝ).smul_right ca) := +(integral_has_fderiv_within_at_of_tendsto_ae hf hmeas_a hmeas_b ha hb).fderiv_within $ hs.prod ht + +/-- Fundamental theorem of calculus: if `f : ℝ → E` is integrable on `a..b` and `f x` has a finite +limit `c` almost surely as `x` tends to `b` from the right or from the left, +then `u ↦ ∫ x in a..u, f x` has right (resp., left) derivative `c` at `b`. -/ +lemma integral_has_deriv_within_at_of_tendsto_ae_right + (hf : interval_integrable f volume a b) {s t : set ℝ} [FTC_filter b (𝓝[s] b) (𝓝[t] b)] + (hmeas : strongly_measurable_at_filter f (𝓝[t] b)) (hb : tendsto f (𝓝[t] b ⊓ volume.ae) (𝓝 c)) : + has_deriv_within_at (λ u, ∫ x in a..u, f x) c s b := +integral_sub_integral_sub_linear_is_o_of_tendsto_ae_right hf hmeas hb + (tendsto_const_pure.mono_right FTC_filter.pure_le) tendsto_id + +/-- Fundamental theorem of calculus: if `f : ℝ → E` is integrable on `a..b` and `f x` is continuous +from the left or from the right at `b`, then `u ↦ ∫ x in a..u, f x` has left (resp., right) +derivative `f b` at `b`. -/ +lemma integral_has_deriv_within_at_right + (hf : interval_integrable f volume a b) {s t : set ℝ} [FTC_filter b (𝓝[s] b) (𝓝[t] b)] + (hmeas : strongly_measurable_at_filter f (𝓝[t] b)) (hb : continuous_within_at f t b) : + has_deriv_within_at (λ u, ∫ x in a..u, f x) (f b) s b := +integral_has_deriv_within_at_of_tendsto_ae_right hf hmeas (hb.mono_left inf_le_left) + +/-- Fundamental theorem of calculus: if `f : ℝ → E` is integrable on `a..b` and `f x` has a finite +limit `c` almost surely as `x` tends to `b` from the right or from the left, then the right +(resp., left) derivative of `u ↦ ∫ x in a..u, f x` at `b` equals `c`. -/ +lemma deriv_within_integral_of_tendsto_ae_right + (hf : interval_integrable f volume a b) {s t : set ℝ} [FTC_filter b (𝓝[s] b) (𝓝[t] b)] + (hmeas: strongly_measurable_at_filter f (𝓝[t] b)) (hb : tendsto f (𝓝[t] b ⊓ volume.ae) (𝓝 c)) + (hs : unique_diff_within_at ℝ s b . unique_diff_within_at_Ici_Iic_univ) : + deriv_within (λ u, ∫ x in a..u, f x) s b = c := +(integral_has_deriv_within_at_of_tendsto_ae_right hf hmeas hb).deriv_within hs + +/-- Fundamental theorem of calculus: if `f : ℝ → E` is integrable on `a..b` and `f x` is continuous +on the right or on the left at `b`, then the right (resp., left) derivative of +`u ↦ ∫ x in a..u, f x` at `b` equals `f b`. -/ +lemma deriv_within_integral_right + (hf : interval_integrable f volume a b) {s t : set ℝ} [FTC_filter b (𝓝[s] b) (𝓝[t] b)] + (hmeas : strongly_measurable_at_filter f (𝓝[t] b)) (hb : continuous_within_at f t b) + (hs : unique_diff_within_at ℝ s b . unique_diff_within_at_Ici_Iic_univ) : + deriv_within (λ u, ∫ x in a..u, f x) s b = f b := +(integral_has_deriv_within_at_right hf hmeas hb).deriv_within hs + +/-- Fundamental theorem of calculus: if `f : ℝ → E` is integrable on `a..b` and `f x` has a finite +limit `c` almost surely as `x` tends to `a` from the right or from the left, +then `u ↦ ∫ x in u..b, f x` has right (resp., left) derivative `-c` at `a`. -/ +lemma integral_has_deriv_within_at_of_tendsto_ae_left + (hf : interval_integrable f volume a b) {s t : set ℝ} [FTC_filter a (𝓝[s] a) (𝓝[t] a)] + (hmeas : strongly_measurable_at_filter f (𝓝[t] a)) + (ha : tendsto f (𝓝[t] a ⊓ volume.ae) (𝓝 c)) : + has_deriv_within_at (λ u, ∫ x in u..b, f x) (-c) s a := +by { simp only [integral_symm b], + exact (integral_has_deriv_within_at_of_tendsto_ae_right hf.symm hmeas ha).neg } + +/-- Fundamental theorem of calculus: if `f : ℝ → E` is integrable on `a..b` and `f x` is continuous +from the left or from the right at `a`, then `u ↦ ∫ x in u..b, f x` has left (resp., right) +derivative `-f a` at `a`. -/ +lemma integral_has_deriv_within_at_left + (hf : interval_integrable f volume a b) {s t : set ℝ} [FTC_filter a (𝓝[s] a) (𝓝[t] a)] + (hmeas : strongly_measurable_at_filter f (𝓝[t] a)) (ha : continuous_within_at f t a) : + has_deriv_within_at (λ u, ∫ x in u..b, f x) (-f a) s a := +integral_has_deriv_within_at_of_tendsto_ae_left hf hmeas (ha.mono_left inf_le_left) + +/-- Fundamental theorem of calculus: if `f : ℝ → E` is integrable on `a..b` and `f x` has a finite +limit `c` almost surely as `x` tends to `a` from the right or from the left, then the right +(resp., left) derivative of `u ↦ ∫ x in u..b, f x` at `a` equals `-c`. -/ +lemma deriv_within_integral_of_tendsto_ae_left + (hf : interval_integrable f volume a b) {s t : set ℝ} [FTC_filter a (𝓝[s] a) (𝓝[t] a)] + (hmeas : strongly_measurable_at_filter f (𝓝[t] a)) (ha : tendsto f (𝓝[t] a ⊓ volume.ae) (𝓝 c)) + (hs : unique_diff_within_at ℝ s a . unique_diff_within_at_Ici_Iic_univ) : + deriv_within (λ u, ∫ x in u..b, f x) s a = -c := +(integral_has_deriv_within_at_of_tendsto_ae_left hf hmeas ha).deriv_within hs + +/-- Fundamental theorem of calculus: if `f : ℝ → E` is integrable on `a..b` and `f x` is continuous +on the right or on the left at `a`, then the right (resp., left) derivative of +`u ↦ ∫ x in u..b, f x` at `a` equals `-f a`. -/ +lemma deriv_within_integral_left + (hf : interval_integrable f volume a b) {s t : set ℝ} [FTC_filter a (𝓝[s] a) (𝓝[t] a)] + (hmeas : strongly_measurable_at_filter f (𝓝[t] a)) (ha : continuous_within_at f t a) + (hs : unique_diff_within_at ℝ s a . unique_diff_within_at_Ici_Iic_univ) : + deriv_within (λ u, ∫ x in u..b, f x) s a = -f a := +(integral_has_deriv_within_at_left hf hmeas ha).deriv_within hs + +/-- The integral of a continuous function is differentiable on a real set `s`. -/ +theorem differentiable_on_integral_of_continuous {s : set ℝ} + (hintg : ∀ x ∈ s, interval_integrable f volume a x) (hcont : continuous f) : + differentiable_on ℝ (λ u, ∫ x in a..u, f x) s := +λ y hy, (integral_has_deriv_at_right (hintg y hy) + hcont.ae_strongly_measurable.strongly_measurable_at_filter + hcont.continuous_at) .differentiable_at.differentiable_within_at + +/-! +### Fundamental theorem of calculus, part 2 + +This section contains theorems pertaining to FTC-2 for interval integrals, i.e., the assertion +that `∫ x in a..b, f' x = f b - f a` under suitable assumptions. + +The most classical version of this theorem assumes that `f'` is continuous. However, this is +unnecessarily strong: the result holds if `f'` is just integrable. We prove the strong version, +following [Rudin, *Real and Complex Analysis* (Theorem 7.21)][rudin2006real]. The proof is first +given for real-valued functions, and then deduced for functions with a general target space. For +a real-valued function `g`, it suffices to show that `g b - g a ≤ (∫ x in a..b, g' x) + ε` for all +positive `ε`. To prove this, choose a lower-semicontinuous function `G'` with `g' < G'` and with +integral close to that of `g'` (its existence is guaranteed by the Vitali-Carathéodory theorem). +It satisfies `g t - g a ≤ ∫ x in a..t, G' x` for all `t ∈ [a, b]`: this inequality holds at `a`, +and if it holds at `t` then it holds for `u` close to `t` on its right, as the left hand side +increases by `g u - g t ∼ (u -t) g' t`, while the right hand side increases by +`∫ x in t..u, G' x` which is roughly at least `∫ x in t..u, G' t = (u - t) G' t`, by lower +semicontinuity. As `g' t < G' t`, this gives the conclusion. One can therefore push progressively +this inequality to the right until the point `b`, where it gives the desired conclusion. +-/ + +variables {g' g φ : ℝ → ℝ} + +/-- Hard part of FTC-2 for integrable derivatives, real-valued functions: one has +`g b - g a ≤ ∫ y in a..b, g' y` when `g'` is integrable. +Auxiliary lemma in the proof of `integral_eq_sub_of_has_deriv_right_of_le`. +We give the slightly more general version that `g b - g a ≤ ∫ y in a..b, φ y` when `g' ≤ φ` and +`φ` is integrable (even if `g'` is not known to be integrable). +Version assuming that `g` is differentiable on `[a, b)`. -/ +lemma sub_le_integral_of_has_deriv_right_of_le_Ico (hab : a ≤ b) (hcont : continuous_on g (Icc a b)) + (hderiv : ∀ x ∈ Ico a b, has_deriv_within_at g (g' x) (Ioi x) x) + (φint : integrable_on φ (Icc a b)) (hφg : ∀ x ∈ Ico a b, g' x ≤ φ x) : + g b - g a ≤ ∫ y in a..b, φ y := +begin + refine le_of_forall_pos_le_add (λ ε εpos, _), + -- Bound from above `g'` by a lower-semicontinuous function `G'`. + rcases exists_lt_lower_semicontinuous_integral_lt φ φint εpos with + ⟨G', f_lt_G', G'cont, G'int, G'lt_top, hG'⟩, + -- we will show by "induction" that `g t - g a ≤ ∫ u in a..t, G' u` for all `t ∈ [a, b]`. + set s := {t | g t - g a ≤ ∫ u in a..t, (G' u).to_real} ∩ Icc a b, + -- the set `s` of points where this property holds is closed. + have s_closed : is_closed s, + { have : continuous_on (λ t, (g t - g a, ∫ u in a..t, (G' u).to_real)) (Icc a b), + { rw ← uIcc_of_le hab at G'int ⊢ hcont, + exact (hcont.sub continuous_on_const).prod (continuous_on_primitive_interval G'int) }, + simp only [s, inter_comm], + exact this.preimage_closed_of_closed is_closed_Icc order_closed_topology.is_closed_le' }, + have main : Icc a b ⊆ {t | g t - g a ≤ ∫ u in a..t, (G' u).to_real }, + { -- to show that the set `s` is all `[a, b]`, it suffices to show that any point `t` in `s` + -- with `t < b` admits another point in `s` slightly to its right + -- (this is a sort of real induction). + apply s_closed.Icc_subset_of_forall_exists_gt + (by simp only [integral_same, mem_set_of_eq, sub_self]) (λ t ht v t_lt_v, _), + obtain ⟨y, g'_lt_y', y_lt_G'⟩ : ∃ (y : ℝ), (g' t : ereal) < y ∧ (y : ereal) < G' t := + ereal.lt_iff_exists_real_btwn.1 ((ereal.coe_le_coe_iff.2 (hφg t ht.2)).trans_lt (f_lt_G' t)), + -- bound from below the increase of `∫ x in a..u, G' x` on the right of `t`, using the lower + -- semicontinuity of `G'`. + have I1 : ∀ᶠ u in 𝓝[>] t, (u - t) * y ≤ ∫ w in t..u, (G' w).to_real, + { have B : ∀ᶠ u in 𝓝 t, (y : ereal) < G' u := + G'cont.lower_semicontinuous_at _ _ y_lt_G', + rcases mem_nhds_iff_exists_Ioo_subset.1 B with ⟨m, M, ⟨hm, hM⟩, H⟩, + have : Ioo t (min M b) ∈ 𝓝[>] t := mem_nhds_within_Ioi_iff_exists_Ioo_subset.2 + ⟨min M b, by simp only [hM, ht.right.right, lt_min_iff, mem_Ioi, and_self], subset.refl _⟩, + filter_upwards [this] with u hu, + have I : Icc t u ⊆ Icc a b := Icc_subset_Icc ht.2.1 (hu.2.le.trans (min_le_right _ _)), + calc (u - t) * y = ∫ v in Icc t u, y : + by simp only [hu.left.le, measure_theory.integral_const, algebra.id.smul_eq_mul, sub_nonneg, + measurable_set.univ, real.volume_Icc, measure.restrict_apply, univ_inter, + ennreal.to_real_of_real] + ... ≤ ∫ w in t..u, (G' w).to_real : + begin + rw [interval_integral.integral_of_le hu.1.le, ← integral_Icc_eq_integral_Ioc], + apply set_integral_mono_ae_restrict, + { simp only [integrable_on_const, real.volume_Icc, ennreal.of_real_lt_top, or_true] }, + { exact integrable_on.mono_set G'int I }, + { have C1 : ∀ᵐ (x : ℝ) ∂volume.restrict (Icc t u), G' x < ∞ := + ae_mono (measure.restrict_mono I le_rfl) G'lt_top, + have C2 : ∀ᵐ (x : ℝ) ∂volume.restrict (Icc t u), x ∈ Icc t u := + ae_restrict_mem measurable_set_Icc, + filter_upwards [C1, C2] with x G'x hx, + apply ereal.coe_le_coe_iff.1, + have : x ∈ Ioo m M, by simp only [hm.trans_le hx.left, + (hx.right.trans_lt hu.right).trans_le (min_le_left M b), mem_Ioo, and_self], + convert le_of_lt (H this), + exact ereal.coe_to_real G'x.ne (ne_bot_of_gt (f_lt_G' x)) } + end }, + -- bound from above the increase of `g u - g a` on the right of `t`, using the derivative at `t` + have I2 : ∀ᶠ u in 𝓝[>] t, g u - g t ≤ (u - t) * y, + { have g'_lt_y : g' t < y := ereal.coe_lt_coe_iff.1 g'_lt_y', + filter_upwards [(hderiv t ⟨ht.2.1, ht.2.2⟩).limsup_slope_le' + (not_mem_Ioi.2 le_rfl) g'_lt_y, self_mem_nhds_within] with u hu t_lt_u, + have := mul_le_mul_of_nonneg_left hu.le (sub_pos.2 t_lt_u).le, + rwa [← smul_eq_mul, sub_smul_slope] at this }, + -- combine the previous two bounds to show that `g u - g a` increases less quickly than + -- `∫ x in a..u, G' x`. + have I3 : ∀ᶠ u in 𝓝[>] t, g u - g t ≤ ∫ w in t..u, (G' w).to_real, + { filter_upwards [I1, I2] with u hu1 hu2 using hu2.trans hu1, }, + have I4 : ∀ᶠ u in 𝓝[>] t, u ∈ Ioc t (min v b), + { refine mem_nhds_within_Ioi_iff_exists_Ioc_subset.2 ⟨min v b, _, subset.refl _⟩, + simp only [lt_min_iff, mem_Ioi], + exact ⟨t_lt_v, ht.2.2⟩ }, + -- choose a point `x` slightly to the right of `t` which satisfies the above bound + rcases (I3.and I4).exists with ⟨x, hx, h'x⟩, + -- we check that it belongs to `s`, essentially by construction + refine ⟨x, _, Ioc_subset_Ioc le_rfl (min_le_left _ _) h'x⟩, + calc g x - g a = (g t - g a) + (g x - g t) : by abel + ... ≤ (∫ w in a..t, (G' w).to_real) + ∫ w in t..x, (G' w).to_real : add_le_add ht.1 hx + ... = ∫ w in a..x, (G' w).to_real : + begin + apply integral_add_adjacent_intervals, + { rw interval_integrable_iff_integrable_Ioc_of_le ht.2.1, + exact integrable_on.mono_set G'int + (Ioc_subset_Icc_self.trans (Icc_subset_Icc le_rfl ht.2.2.le)) }, + { rw interval_integrable_iff_integrable_Ioc_of_le h'x.1.le, + apply integrable_on.mono_set G'int, + refine Ioc_subset_Icc_self.trans (Icc_subset_Icc ht.2.1 (h'x.2.trans (min_le_right _ _))) } + end }, + -- now that we know that `s` contains `[a, b]`, we get the desired result by applying this to `b`. + calc g b - g a ≤ ∫ y in a..b, (G' y).to_real : main (right_mem_Icc.2 hab) + ... ≤ (∫ y in a..b, φ y) + ε : + begin + convert hG'.le; + { rw interval_integral.integral_of_le hab, + simp only [integral_Icc_eq_integral_Ioc', real.volume_singleton] }, + end +end + +/-- Hard part of FTC-2 for integrable derivatives, real-valued functions: one has +`g b - g a ≤ ∫ y in a..b, g' y` when `g'` is integrable. +Auxiliary lemma in the proof of `integral_eq_sub_of_has_deriv_right_of_le`. +We give the slightly more general version that `g b - g a ≤ ∫ y in a..b, φ y` when `g' ≤ φ` and +`φ` is integrable (even if `g'` is not known to be integrable). +Version assuming that `g` is differentiable on `(a, b)`. -/ +lemma sub_le_integral_of_has_deriv_right_of_le (hab : a ≤ b) + (hcont : continuous_on g (Icc a b)) + (hderiv : ∀ x ∈ Ioo a b, has_deriv_within_at g (g' x) (Ioi x) x) + (φint : integrable_on φ (Icc a b)) (hφg : ∀ x ∈ Ioo a b, g' x ≤ φ x) : + g b - g a ≤ ∫ y in a..b, φ y := +begin + -- This follows from the version on a closed-open interval (applied to `[t, b)` for `t` close to + -- `a`) and a continuity argument. + obtain rfl|a_lt_b := hab.eq_or_lt, { simp }, + set s := {t | g b - g t ≤ ∫ u in t..b, φ u} ∩ Icc a b, + have s_closed : is_closed s, + { have : continuous_on (λ t, (g b - g t, ∫ u in t..b, φ u)) (Icc a b), + { rw ← uIcc_of_le hab at ⊢ hcont φint, + exact (continuous_on_const.sub hcont).prod (continuous_on_primitive_interval_left φint) }, + simp only [s, inter_comm], + exact this.preimage_closed_of_closed is_closed_Icc is_closed_le_prod, }, + have A : closure (Ioc a b) ⊆ s, + { apply s_closed.closure_subset_iff.2, + assume t ht, + refine ⟨_, ⟨ht.1.le, ht.2⟩⟩, + exact sub_le_integral_of_has_deriv_right_of_le_Ico ht.2 + (hcont.mono (Icc_subset_Icc ht.1.le le_rfl)) + (λ x hx, hderiv x ⟨ht.1.trans_le hx.1, hx.2⟩) + (φint.mono_set (Icc_subset_Icc ht.1.le le_rfl)) + (λ x hx, hφg x ⟨ht.1.trans_le hx.1, hx.2⟩) }, + rw closure_Ioc a_lt_b.ne at A, + exact (A (left_mem_Icc.2 hab)).1, +end + +/-- Auxiliary lemma in the proof of `integral_eq_sub_of_has_deriv_right_of_le`. -/ +lemma integral_le_sub_of_has_deriv_right_of_le (hab : a ≤ b) + (hcont : continuous_on g (Icc a b)) + (hderiv : ∀ x ∈ Ioo a b, has_deriv_within_at g (g' x) (Ioi x) x) + (φint : integrable_on φ (Icc a b)) (hφg : ∀ x ∈ Ioo a b, φ x ≤ g' x) : + ∫ y in a..b, φ y ≤ g b - g a := +begin + rw ← neg_le_neg_iff, + convert sub_le_integral_of_has_deriv_right_of_le hab hcont.neg (λ x hx, (hderiv x hx).neg) + φint.neg (λ x hx, neg_le_neg (hφg x hx)), + { abel }, + { simp only [← integral_neg], refl }, +end + +/-- Auxiliary lemma in the proof of `integral_eq_sub_of_has_deriv_right_of_le`: real version -/ +lemma integral_eq_sub_of_has_deriv_right_of_le_real (hab : a ≤ b) + (hcont : continuous_on g (Icc a b)) + (hderiv : ∀ x ∈ Ioo a b, has_deriv_within_at g (g' x) (Ioi x) x) + (g'int : integrable_on g' (Icc a b)) : + ∫ y in a..b, g' y = g b - g a := +le_antisymm + (integral_le_sub_of_has_deriv_right_of_le hab hcont hderiv g'int (λ x hx, le_rfl)) + (sub_le_integral_of_has_deriv_right_of_le hab hcont hderiv g'int (λ x hx, le_rfl)) + +variable {f' : ℝ → E} + +/-- **Fundamental theorem of calculus-2**: If `f : ℝ → E` is continuous on `[a, b]` (where `a ≤ b`) + and has a right derivative at `f' x` for all `x` in `(a, b)`, and `f'` is integrable on `[a, b]`, + then `∫ y in a..b, f' y` equals `f b - f a`. -/ +theorem integral_eq_sub_of_has_deriv_right_of_le (hab : a ≤ b) (hcont : continuous_on f (Icc a b)) + (hderiv : ∀ x ∈ Ioo a b, has_deriv_within_at f (f' x) (Ioi x) x) + (f'int : interval_integrable f' volume a b) : + ∫ y in a..b, f' y = f b - f a := +begin + refine (normed_space.eq_iff_forall_dual_eq ℝ).2 (λ g, _), + rw [← g.interval_integral_comp_comm f'int, g.map_sub], + exact integral_eq_sub_of_has_deriv_right_of_le_real hab (g.continuous.comp_continuous_on hcont) + (λ x hx, g.has_fderiv_at.comp_has_deriv_within_at x (hderiv x hx)) + (g.integrable_comp ((interval_integrable_iff_integrable_Icc_of_le hab).1 f'int)) +end + +/-- Fundamental theorem of calculus-2: If `f : ℝ → E` is continuous on `[a, b]` and + has a right derivative at `f' x` for all `x` in `[a, b)`, and `f'` is integrable on `[a, b]` then + `∫ y in a..b, f' y` equals `f b - f a`. -/ +theorem integral_eq_sub_of_has_deriv_right (hcont : continuous_on f (uIcc a b)) + (hderiv : ∀ x ∈ Ioo (min a b) (max a b), has_deriv_within_at f (f' x) (Ioi x) x) + (hint : interval_integrable f' volume a b) : + ∫ y in a..b, f' y = f b - f a := +begin + cases le_total a b with hab hab, + { simp only [uIcc_of_le, min_eq_left, max_eq_right, hab] at hcont hderiv hint, + apply integral_eq_sub_of_has_deriv_right_of_le hab hcont hderiv hint }, + { simp only [uIcc_of_ge, min_eq_right, max_eq_left, hab] at hcont hderiv, + rw [integral_symm, integral_eq_sub_of_has_deriv_right_of_le hab hcont hderiv hint.symm, + neg_sub] } +end + +/-- Fundamental theorem of calculus-2: If `f : ℝ → E` is continuous on `[a, b]` (where `a ≤ b`) and + has a derivative at `f' x` for all `x` in `(a, b)`, and `f'` is integrable on `[a, b]`, then + `∫ y in a..b, f' y` equals `f b - f a`. -/ +theorem integral_eq_sub_of_has_deriv_at_of_le (hab : a ≤ b) + (hcont : continuous_on f (Icc a b)) + (hderiv : ∀ x ∈ Ioo a b, has_deriv_at f (f' x) x) (hint : interval_integrable f' volume a b) : + ∫ y in a..b, f' y = f b - f a := +integral_eq_sub_of_has_deriv_right_of_le hab hcont (λ x hx, (hderiv x hx).has_deriv_within_at) hint + +/-- Fundamental theorem of calculus-2: If `f : ℝ → E` has a derivative at `f' x` for all `x` in + `[a, b]` and `f'` is integrable on `[a, b]`, then `∫ y in a..b, f' y` equals `f b - f a`. -/ +theorem integral_eq_sub_of_has_deriv_at + (hderiv : ∀ x ∈ uIcc a b, has_deriv_at f (f' x) x) + (hint : interval_integrable f' volume a b) : + ∫ y in a..b, f' y = f b - f a := +integral_eq_sub_of_has_deriv_right (has_deriv_at.continuous_on hderiv) + (λ x hx, (hderiv _ (mem_Icc_of_Ioo hx)).has_deriv_within_at) hint + +theorem integral_eq_sub_of_has_deriv_at_of_tendsto (hab : a < b) {fa fb} + (hderiv : ∀ x ∈ Ioo a b, has_deriv_at f (f' x) x) (hint : interval_integrable f' volume a b) + (ha : tendsto f (𝓝[>] a) (𝓝 fa)) (hb : tendsto f (𝓝[<] b) (𝓝 fb)) : + ∫ y in a..b, f' y = fb - fa := +begin + set F : ℝ → E := update (update f a fa) b fb, + have Fderiv : ∀ x ∈ Ioo a b, has_deriv_at F (f' x) x, + { refine λ x hx, (hderiv x hx).congr_of_eventually_eq _, + filter_upwards [Ioo_mem_nhds hx.1 hx.2] with _ hy, simp only [F], + rw [update_noteq hy.2.ne, update_noteq hy.1.ne'], }, + have hcont : continuous_on F (Icc a b), + { rw [continuous_on_update_iff, continuous_on_update_iff, Icc_diff_right, Ico_diff_left], + refine ⟨⟨λ z hz, (hderiv z hz).continuous_at.continuous_within_at, _⟩, _⟩, + { exact λ _, ha.mono_left (nhds_within_mono _ Ioo_subset_Ioi_self) }, + { rintro -, + refine (hb.congr' _).mono_left (nhds_within_mono _ Ico_subset_Iio_self), + filter_upwards [Ioo_mem_nhds_within_Iio (right_mem_Ioc.2 hab)] + with _ hz using (update_noteq hz.1.ne' _ _).symm } }, + simpa [F, hab.ne, hab.ne'] using integral_eq_sub_of_has_deriv_at_of_le hab.le hcont Fderiv hint +end + +/-- Fundamental theorem of calculus-2: If `f : ℝ → E` is differentiable at every `x` in `[a, b]` and + its derivative is integrable on `[a, b]`, then `∫ y in a..b, deriv f y` equals `f b - f a`. -/ +theorem integral_deriv_eq_sub (hderiv : ∀ x ∈ uIcc a b, differentiable_at ℝ f x) + (hint : interval_integrable (deriv f) volume a b) : + ∫ y in a..b, deriv f y = f b - f a := +integral_eq_sub_of_has_deriv_at (λ x hx, (hderiv x hx).has_deriv_at) hint + +theorem integral_deriv_eq_sub' (f) (hderiv : deriv f = f') + (hdiff : ∀ x ∈ uIcc a b, differentiable_at ℝ f x) + (hcont : continuous_on f' (uIcc a b)) : + ∫ y in a..b, f' y = f b - f a := +begin + rw [← hderiv, integral_deriv_eq_sub hdiff], + rw hderiv, + exact hcont.interval_integrable +end + +/-! +### Automatic integrability for nonnegative derivatives +-/ + +/-- When the right derivative of a function is nonnegative, then it is automatically integrable. -/ +lemma integrable_on_deriv_right_of_nonneg (hcont : continuous_on g (Icc a b)) + (hderiv : ∀ x ∈ Ioo a b, has_deriv_within_at g (g' x) (Ioi x) x) + (g'pos : ∀ x ∈ Ioo a b, 0 ≤ g' x) : + integrable_on g' (Ioc a b) := +begin + by_cases hab : a < b, swap, + { simp [Ioc_eq_empty hab] }, + rw integrable_on_Ioc_iff_integrable_on_Ioo, + have meas_g' : ae_measurable g' (volume.restrict (Ioo a b)), + { apply (ae_measurable_deriv_within_Ioi g _).congr, + refine (ae_restrict_mem measurable_set_Ioo).mono (λ x hx, _), + exact (hderiv x hx).deriv_within (unique_diff_within_at_Ioi _) }, + suffices H : ∫⁻ x in Ioo a b, ‖g' x‖₊ ≤ ennreal.of_real (g b - g a), + from ⟨meas_g'.ae_strongly_measurable, H.trans_lt ennreal.of_real_lt_top⟩, + by_contra' H, + obtain ⟨f, fle, fint, hf⟩ : + ∃ (f : simple_func ℝ ℝ≥0), (∀ x, f x ≤ ‖g' x‖₊) ∧ ∫⁻ (x : ℝ) in Ioo a b, f x < ∞ + ∧ ennreal.of_real (g b - g a) < ∫⁻ (x : ℝ) in Ioo a b, f x := + exists_lt_lintegral_simple_func_of_lt_lintegral H, + let F : ℝ → ℝ := coe ∘ f, + have intF : integrable_on F (Ioo a b), + { refine ⟨f.measurable.coe_nnreal_real.ae_strongly_measurable, _⟩, + simpa only [has_finite_integral, nnreal.nnnorm_eq] using fint }, + have A : ∫⁻ (x : ℝ) in Ioo a b, f x = ennreal.of_real (∫ x in Ioo a b, F x) := + lintegral_coe_eq_integral _ intF, + rw A at hf, + have B : ∫ (x : ℝ) in Ioo a b, F x ≤ g b - g a, + { rw [← integral_Ioc_eq_integral_Ioo, ← interval_integral.integral_of_le hab.le], + apply integral_le_sub_of_has_deriv_right_of_le hab.le hcont hderiv _ (λ x hx, _), + { rwa integrable_on_Icc_iff_integrable_on_Ioo }, + { convert nnreal.coe_le_coe.2 (fle x), + simp only [real.norm_of_nonneg (g'pos x hx), coe_nnnorm] } }, + exact lt_irrefl _ (hf.trans_le (ennreal.of_real_le_of_real B)), +end + +/-- When the derivative of a function is nonnegative, then it is automatically integrable, +Ioc version. -/ +lemma integrable_on_deriv_of_nonneg (hcont : continuous_on g (Icc a b)) + (hderiv : ∀ x ∈ Ioo a b, has_deriv_at g (g' x) x) + (g'pos : ∀ x ∈ Ioo a b, 0 ≤ g' x) : + integrable_on g' (Ioc a b) := +integrable_on_deriv_right_of_nonneg hcont (λ x hx, (hderiv x hx).has_deriv_within_at) g'pos + +/-- When the derivative of a function is nonnegative, then it is automatically integrable, +interval version. -/ +theorem interval_integrable_deriv_of_nonneg (hcont : continuous_on g (uIcc a b)) + (hderiv : ∀ x ∈ Ioo (min a b) (max a b), has_deriv_at g (g' x) x) + (hpos : ∀ x ∈ Ioo (min a b) (max a b), 0 ≤ g' x) : + interval_integrable g' volume a b := +begin + cases le_total a b with hab hab, + { simp only [uIcc_of_le, min_eq_left, max_eq_right, hab, interval_integrable, + hab, Ioc_eq_empty_of_le, integrable_on_empty, and_true] at hcont hderiv hpos ⊢, + exact integrable_on_deriv_of_nonneg hcont hderiv hpos, }, + { simp only [uIcc_of_ge, min_eq_right, max_eq_left, hab, interval_integrable, + Ioc_eq_empty_of_le, integrable_on_empty, true_and] at hcont hderiv hpos ⊢, + exact integrable_on_deriv_of_nonneg hcont hderiv hpos } +end + +/-! +### Integration by parts +-/ + +section parts + +variables [normed_ring A] [normed_algebra ℝ A] [complete_space A] + +theorem integral_deriv_mul_eq_sub {u v u' v' : ℝ → A} + (hu : ∀ x ∈ uIcc a b, has_deriv_at u (u' x) x) + (hv : ∀ x ∈ uIcc a b, has_deriv_at v (v' x) x) + (hu' : interval_integrable u' volume a b) (hv' : interval_integrable v' volume a b) : + ∫ x in a..b, u' x * v x + u x * v' x = u b * v b - u a * v a := +integral_eq_sub_of_has_deriv_at (λ x hx, (hu x hx).mul (hv x hx)) $ + (hu'.mul_continuous_on (has_deriv_at.continuous_on hv)).add + (hv'.continuous_on_mul ((has_deriv_at.continuous_on hu))) + +theorem integral_mul_deriv_eq_deriv_mul {u v u' v' : ℝ → A} + (hu : ∀ x ∈ uIcc a b, has_deriv_at u (u' x) x) + (hv : ∀ x ∈ uIcc a b, has_deriv_at v (v' x) x) + (hu' : interval_integrable u' volume a b) (hv' : interval_integrable v' volume a b) : + ∫ x in a..b, u x * v' x = u b * v b - u a * v a - ∫ x in a..b, u' x * v x := +begin + rw [← integral_deriv_mul_eq_sub hu hv hu' hv', ← integral_sub], + { exact integral_congr (λ x hx, by simp only [add_sub_cancel']) }, + { exact ((hu'.mul_continuous_on (has_deriv_at.continuous_on hv)).add + (hv'.continuous_on_mul (has_deriv_at.continuous_on hu))) }, + { exact hu'.mul_continuous_on (has_deriv_at.continuous_on hv) }, +end + +end parts + +/-! +### Integration by substitution / Change of variables +-/ + +section smul + +/-- +Change of variables, general form. If `f` is continuous on `[a, b]` and has +right-derivative `f'` in `(a, b)`, `g` is continuous on `f '' (a, b)` and integrable on +`f '' [a, b]`, and `f' x • (g ∘ f) x` is integrable on `[a, b]`, +then we can substitute `u = f x` to get `∫ x in a..b, f' x • (g ∘ f) x = ∫ u in f a..f b, g u`. +-/ +theorem integral_comp_smul_deriv''' {f f' : ℝ → ℝ} {g : ℝ → E} + (hf : continuous_on f [a, b]) + (hff' : ∀ x ∈ Ioo (min a b) (max a b), has_deriv_within_at f (f' x) (Ioi x) x) + (hg_cont : continuous_on g (f '' Ioo (min a b) (max a b))) + (hg1 : integrable_on g (f '' [a, b]) ) + (hg2 : integrable_on (λ x, f'(x) • (g ∘ f) x) [a, b]) : + ∫ x in a..b, f' x • (g ∘ f) x = ∫ u in f a..f b, g u := +begin + rw [hf.image_uIcc, ←interval_integrable_iff'] at hg1, + have h_cont : continuous_on (λ u, ∫ t in f a..f u, g t) [a, b], + { refine (continuous_on_primitive_interval' hg1 _).comp hf _, + { rw ← hf.image_uIcc, exact mem_image_of_mem f left_mem_uIcc }, + { rw ← hf.image_uIcc, exact maps_to_image _ _ } }, + have h_der : ∀ x ∈ Ioo (min a b) (max a b), has_deriv_within_at + (λ u, ∫ t in f a..f u, g t) (f' x • ((g ∘ f) x)) (Ioi x) x, + { intros x hx, + obtain ⟨c, hc⟩ := nonempty_Ioo.mpr hx.1, + obtain ⟨d, hd⟩ := nonempty_Ioo.mpr hx.2, + have cdsub : [c, d] ⊆ Ioo (min a b) (max a b), + { rw uIcc_of_le (hc.2.trans hd.1).le, exact Icc_subset_Ioo hc.1 hd.2 }, + replace hg_cont := hg_cont.mono (image_subset f cdsub), + let J := [Inf (f '' [c, d]), Sup (f '' [c, d])], + have hJ : f '' [c, d] = J := (hf.mono (cdsub.trans Ioo_subset_Icc_self)).image_uIcc, + rw hJ at hg_cont, + have h2x : f x ∈ J, { rw ←hJ, exact mem_image_of_mem _ (mem_uIcc_of_le hc.2.le hd.1.le), }, + have h2g : interval_integrable g volume (f a) (f x), + { refine hg1.mono_set _, + rw ←hf.image_uIcc, + exact hf.surj_on_uIcc left_mem_uIcc (Ioo_subset_Icc_self hx) }, + have h3g := hg_cont.strongly_measurable_at_filter_nhds_within measurable_set_Icc (f x), + haveI : fact (f x ∈ J) := ⟨h2x⟩, + have : has_deriv_within_at (λ u, ∫ x in f a..u, g x) (g (f x)) J (f x) := + interval_integral.integral_has_deriv_within_at_right h2g h3g (hg_cont (f x) h2x), + refine (this.scomp x ((hff' x hx).Ioo_of_Ioi hd.1) _).Ioi_of_Ioo hd.1, + rw ←hJ, + refine (maps_to_image _ _).mono _ subset.rfl, + exact Ioo_subset_Icc_self.trans ((Icc_subset_Icc_left hc.2.le).trans Icc_subset_uIcc) }, + rw ←interval_integrable_iff' at hg2, + simp_rw [integral_eq_sub_of_has_deriv_right h_cont h_der hg2, integral_same, sub_zero], +end + +/-- +Change of variables for continuous integrands. If `f` is continuous on `[a, b]` and has +continuous right-derivative `f'` in `(a, b)`, and `g` is continuous on `f '' [a, b]` then we can +substitute `u = f x` to get `∫ x in a..b, f' x • (g ∘ f) x = ∫ u in f a..f b, g u`. +-/ +theorem integral_comp_smul_deriv'' {f f' : ℝ → ℝ} {g : ℝ → E} + (hf : continuous_on f [a, b]) + (hff' : ∀ x ∈ Ioo (min a b) (max a b), has_deriv_within_at f (f' x) (Ioi x) x) + (hf' : continuous_on f' [a, b]) + (hg : continuous_on g (f '' [a, b])) : + ∫ x in a..b, f' x • (g ∘ f) x = ∫ u in f a..f b, g u := +begin + refine integral_comp_smul_deriv''' hf hff' + (hg.mono $ image_subset _ Ioo_subset_Icc_self) _ + (hf'.smul (hg.comp hf $ subset_preimage_image f _)).integrable_on_Icc, + rw hf.image_uIcc at hg ⊢, + exact hg.integrable_on_Icc, +end + +/-- +Change of variables. If `f` is has continuous derivative `f'` on `[a, b]`, +and `g` is continuous on `f '' [a, b]`, then we can substitute `u = f x` to get +`∫ x in a..b, f' x • (g ∘ f) x = ∫ u in f a..f b, g u`. +Compared to `interval_integral.integral_comp_smul_deriv` we only require that `g` is continuous on +`f '' [a, b]`. +-/ +theorem integral_comp_smul_deriv' {f f' : ℝ → ℝ} {g : ℝ → E} + (h : ∀ x ∈ uIcc a b, has_deriv_at f (f' x) x) + (h' : continuous_on f' (uIcc a b)) (hg : continuous_on g (f '' [a, b])) : + ∫ x in a..b, f' x • (g ∘ f) x = ∫ x in f a..f b, g x := +integral_comp_smul_deriv'' (λ x hx, (h x hx).continuous_at.continuous_within_at) + (λ x hx, (h x $ Ioo_subset_Icc_self hx).has_deriv_within_at) h' hg + +/-- +Change of variables, most common version. If `f` is has continuous derivative `f'` on `[a, b]`, +and `g` is continuous, then we can substitute `u = f x` to get +`∫ x in a..b, f' x • (g ∘ f) x = ∫ u in f a..f b, g u`. +-/ +theorem integral_comp_smul_deriv {f f' : ℝ → ℝ} {g : ℝ → E} + (h : ∀ x ∈ uIcc a b, has_deriv_at f (f' x) x) + (h' : continuous_on f' (uIcc a b)) (hg : continuous g) : + ∫ x in a..b, f' x • (g ∘ f) x = ∫ x in f a..f b, g x := +integral_comp_smul_deriv' h h' hg.continuous_on + +theorem integral_deriv_comp_smul_deriv' {f f' : ℝ → ℝ} {g g' : ℝ → E} + (hf : continuous_on f [a, b]) + (hff' : ∀ x ∈ Ioo (min a b) (max a b), has_deriv_within_at f (f' x) (Ioi x) x) + (hf' : continuous_on f' [a, b]) + (hg : continuous_on g [f a, f b]) + (hgg' : ∀ x ∈ Ioo (min (f a) (f b)) (max (f a) (f b)), has_deriv_within_at g (g' x) (Ioi x) x) + (hg' : continuous_on g' (f '' [a, b])) : + ∫ x in a..b, f' x • (g' ∘ f) x = (g ∘ f) b - (g ∘ f) a := +begin + rw [integral_comp_smul_deriv'' hf hff' hf' hg', + integral_eq_sub_of_has_deriv_right hg hgg' (hg'.mono _).interval_integrable], + exact intermediate_value_uIcc hf +end + +theorem integral_deriv_comp_smul_deriv {f f' : ℝ → ℝ} {g g' : ℝ → E} + (hf : ∀ x ∈ uIcc a b, has_deriv_at f (f' x) x) + (hg : ∀ x ∈ uIcc a b, has_deriv_at g (g' (f x)) (f x)) + (hf' : continuous_on f' (uIcc a b)) (hg' : continuous g') : + ∫ x in a..b, f' x • (g' ∘ f) x = (g ∘ f) b - (g ∘ f) a := +integral_eq_sub_of_has_deriv_at (λ x hx, (hg x hx).scomp x $ hf x hx) + (hf'.smul (hg'.comp_continuous_on $ has_deriv_at.continuous_on hf)).interval_integrable + +end smul +section mul + +/-- +Change of variables, general form for scalar functions. If `f` is continuous on `[a, b]` and has +continuous right-derivative `f'` in `(a, b)`, `g` is continuous on `f '' (a, b)` and integrable on +`f '' [a, b]`, and `(g ∘ f) x * f' x` is integrable on `[a, b]`, then we can substitute `u = f x` +to get `∫ x in a..b, (g ∘ f) x * f' x = ∫ u in f a..f b, g u`. +-/ +theorem integral_comp_mul_deriv''' {a b : ℝ} {f f' : ℝ → ℝ} {g : ℝ → ℝ} + (hf : continuous_on f [a, b]) + (hff' : ∀ x ∈ Ioo (min a b) (max a b), has_deriv_within_at f (f' x) (Ioi x) x) + (hg_cont : continuous_on g (f '' Ioo (min a b) (max a b))) + (hg1 : integrable_on g (f '' [a, b]) ) + (hg2 : integrable_on (λ x, (g ∘ f) x * f' x) [a, b]) : + ∫ x in a..b, (g ∘ f) x * f' x = ∫ u in f a..f b, g u := +begin + have hg2' : integrable_on (λ x, f' x • (g ∘ f) x) [a, b] := by simpa [mul_comm] using hg2, + simpa [mul_comm] using integral_comp_smul_deriv''' hf hff' hg_cont hg1 hg2', +end + +/-- +Change of variables for continuous integrands. If `f` is continuous on `[a, b]` and has +continuous right-derivative `f'` in `(a, b)`, and `g` is continuous on `f '' [a, b]` then we can +substitute `u = f x` to get `∫ x in a..b, (g ∘ f) x * f' x = ∫ u in f a..f b, g u`. +-/ +theorem integral_comp_mul_deriv'' {f f' g : ℝ → ℝ} + (hf : continuous_on f [a, b]) + (hff' : ∀ x ∈ Ioo (min a b) (max a b), has_deriv_within_at f (f' x) (Ioi x) x) + (hf' : continuous_on f' [a, b]) + (hg : continuous_on g (f '' [a, b])) : + ∫ x in a..b, (g ∘ f) x * f' x = ∫ u in f a..f b, g u := +by simpa [mul_comm] using integral_comp_smul_deriv'' hf hff' hf' hg + +/-- +Change of variables. If `f` is has continuous derivative `f'` on `[a, b]`, +and `g` is continuous on `f '' [a, b]`, then we can substitute `u = f x` to get +`∫ x in a..b, (g ∘ f) x * f' x = ∫ u in f a..f b, g u`. +Compared to `interval_integral.integral_comp_mul_deriv` we only require that `g` is continuous on +`f '' [a, b]`. +-/ +theorem integral_comp_mul_deriv' {f f' g : ℝ → ℝ} + (h : ∀ x ∈ uIcc a b, has_deriv_at f (f' x) x) + (h' : continuous_on f' (uIcc a b)) (hg : continuous_on g (f '' [a, b])) : + ∫ x in a..b, (g ∘ f) x * f' x = ∫ x in f a..f b, g x := +by simpa [mul_comm] using integral_comp_smul_deriv' h h' hg + +/-- +Change of variables, most common version. If `f` is has continuous derivative `f'` on `[a, b]`, +and `g` is continuous, then we can substitute `u = f x` to get +`∫ x in a..b, (g ∘ f) x * f' x = ∫ u in f a..f b, g u`. +-/ +theorem integral_comp_mul_deriv {f f' g : ℝ → ℝ} + (h : ∀ x ∈ uIcc a b, has_deriv_at f (f' x) x) + (h' : continuous_on f' (uIcc a b)) (hg : continuous g) : + ∫ x in a..b, (g ∘ f) x * f' x = ∫ x in f a..f b, g x := +integral_comp_mul_deriv' h h' hg.continuous_on + +theorem integral_deriv_comp_mul_deriv' {f f' g g' : ℝ → ℝ} + (hf : continuous_on f [a, b]) + (hff' : ∀ x ∈ Ioo (min a b) (max a b), has_deriv_within_at f (f' x) (Ioi x) x) + (hf' : continuous_on f' [a, b]) + (hg : continuous_on g [f a, f b]) + (hgg' : ∀ x ∈ Ioo (min (f a) (f b)) (max (f a) (f b)), has_deriv_within_at g (g' x) (Ioi x) x) + (hg' : continuous_on g' (f '' [a, b])) : + ∫ x in a..b, (g' ∘ f) x * f' x = (g ∘ f) b - (g ∘ f) a := +by simpa [mul_comm] using integral_deriv_comp_smul_deriv' hf hff' hf' hg hgg' hg' + +theorem integral_deriv_comp_mul_deriv {f f' g g' : ℝ → ℝ} + (hf : ∀ x ∈ uIcc a b, has_deriv_at f (f' x) x) + (hg : ∀ x ∈ uIcc a b, has_deriv_at g (g' (f x)) (f x)) + (hf' : continuous_on f' (uIcc a b)) (hg' : continuous g') : + ∫ x in a..b, (g' ∘ f) x * f' x = (g ∘ f) b - (g ∘ f) a := +by simpa [mul_comm] using integral_deriv_comp_smul_deriv hf hg hf' hg' + +end mul + +end interval_integral diff --git a/src/measure_theory/integral/integral_eq_improper.lean b/src/measure_theory/integral/integral_eq_improper.lean index 335574962c38d..83d5b5ce1cfdb 100644 --- a/src/measure_theory/integral/integral_eq_improper.lean +++ b/src/measure_theory/integral/integral_eq_improper.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Anatole Dedecker. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Anatole Dedecker, Bhavik Mehta -/ -import measure_theory.integral.interval_integral +import measure_theory.integral.fund_thm_calculus import order.filter.at_top_bot import measure_theory.function.jacobian diff --git a/src/measure_theory/integral/interval_integral.lean b/src/measure_theory/integral/interval_integral.lean index a9addd4c9b351..e6fb9c5c0a3db 100644 --- a/src/measure_theory/integral/interval_integral.lean +++ b/src/measure_theory/integral/interval_integral.lean @@ -3,102 +3,15 @@ Copyright (c) 2020 Yury G. Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury G. Kudryashov, Patrick Massot, Sébastien Gouëzel -/ -import analysis.normed_space.dual import data.set.intervals.disjoint -import measure_theory.measure.haar_lebesgue -import measure_theory.function.locally_integrable import measure_theory.integral.set_integral -import measure_theory.integral.vitali_caratheodory -import analysis.calculus.fderiv_measurable +import measure_theory.measure.haar_lebesgue /-! # Integral over an interval In this file we define `∫ x in a..b, f x ∂μ` to be `∫ x in Ioc a b, f x ∂μ` if `a ≤ b` and -`-∫ x in Ioc b a, f x ∂μ` if `b ≤ a`. We prove a few simple properties and several versions of the -[fundamental theorem of calculus](https://en.wikipedia.org/wiki/Fundamental_theorem_of_calculus). - -Recall that its first version states that the function `(u, v) ↦ ∫ x in u..v, f x` has derivative -`(δu, δv) ↦ δv • f b - δu • f a` at `(a, b)` provided that `f` is continuous at `a` and `b`, -and its second version states that, if `f` has an integrable derivative on `[a, b]`, then -`∫ x in a..b, f' x = f b - f a`. - -## Main statements - -### FTC-1 for Lebesgue measure - -We prove several versions of FTC-1, all in the `interval_integral` namespace. Many of them follow -the naming scheme `integral_has(_strict?)_(f?)deriv(_within?)_at(_of_tendsto_ae?)(_right|_left?)`. -They formulate FTC in terms of `has(_strict?)_(f?)deriv(_within?)_at`. -Let us explain the meaning of each part of the name: - -* `_strict` means that the theorem is about strict differentiability; -* `f` means that the theorem is about differentiability in both endpoints; incompatible with - `_right|_left`; -* `_within` means that the theorem is about one-sided derivatives, see below for details; -* `_of_tendsto_ae` means that instead of continuity the theorem assumes that `f` has a finite limit - almost surely as `x` tends to `a` and/or `b`; -* `_right` or `_left` mean that the theorem is about differentiability in the right (resp., left) - endpoint. - -We also reformulate these theorems in terms of `(f?)deriv(_within?)`. These theorems are named -`(f?)deriv(_within?)_integral(_of_tendsto_ae?)(_right|_left?)` with the same meaning of parts of the -name. - -### One-sided derivatives - -Theorem `integral_has_fderiv_within_at_of_tendsto_ae` states that `(u, v) ↦ ∫ x in u..v, f x` has a -derivative `(δu, δv) ↦ δv • cb - δu • ca` within the set `s × t` at `(a, b)` provided that `f` tends -to `ca` (resp., `cb`) almost surely at `la` (resp., `lb`), where possible values of `s`, `t`, and -corresponding filters `la`, `lb` are given in the following table. - -| `s` | `la` | `t` | `lb` | -| ------- | ---- | --- | ---- | -| `Iic a` | `𝓝[≤] a` | `Iic b` | `𝓝[≤] b` | -| `Ici a` | `𝓝[>] a` | `Ici b` | `𝓝[>] b` | -| `{a}` | `⊥` | `{b}` | `⊥` | -| `univ` | `𝓝 a` | `univ` | `𝓝 b` | - -We use a typeclass `FTC_filter` to make Lean automatically find `la`/`lb` based on `s`/`t`. This way -we can formulate one theorem instead of `16` (or `8` if we leave only non-trivial ones not covered -by `integral_has_deriv_within_at_of_tendsto_ae_(left|right)` and -`integral_has_fderiv_at_of_tendsto_ae`). Similarly, -`integral_has_deriv_within_at_of_tendsto_ae_right` works for both one-sided derivatives using the -same typeclass to find an appropriate filter. - -### FTC for a locally finite measure - -Before proving FTC for the Lebesgue measure, we prove a few statements that can be seen as FTC for -any measure. The most general of them, -`measure_integral_sub_integral_sub_linear_is_o_of_tendsto_ae`, states the following. Let `(la, la')` -be an `FTC_filter` pair of filters around `a` (i.e., `FTC_filter a la la'`) and let `(lb, lb')` be -an `FTC_filter` pair of filters around `b`. If `f` has finite limits `ca` and `cb` almost surely at -`la'` and `lb'`, respectively, then -`∫ x in va..vb, f x ∂μ - ∫ x in ua..ub, f x ∂μ = ∫ x in ub..vb, cb ∂μ - ∫ x in ua..va, ca ∂μ + - o(‖∫ x in ua..va, (1:ℝ) ∂μ‖ + ‖∫ x in ub..vb, (1:ℝ) ∂μ‖)` as `ua` and `va` tend to `la` while -`ub` and `vb` tend to `lb`. - -### FTC-2 and corollaries - -We use FTC-1 to prove several versions of FTC-2 for the Lebesgue measure, using a similar naming -scheme as for the versions of FTC-1. They include: -* `interval_integral.integral_eq_sub_of_has_deriv_right_of_le` - most general version, for functions - with a right derivative -* `interval_integral.integral_eq_sub_of_has_deriv_at'` - version for functions with a derivative on - an open set -* `interval_integral.integral_deriv_eq_sub'` - version that is easiest to use when computing the - integral of a specific function - -We then derive additional integration techniques from FTC-2: -* `interval_integral.integral_mul_deriv_eq_deriv_mul` - integration by parts -* `interval_integral.integral_comp_mul_deriv''` - integration by substitution - -Many applications of these theorems can be found in the file `analysis.special_functions.integrals`. - -Note that the assumptions of FTC-2 are formulated in the form that `f'` is integrable. To use it in -a context with the stronger assumption that `f'` is continuous, one can use -`continuous_on.interval_integrable` or `continuous_on.integrable_on_Icc` or -`continuous_on.integrable_on_interval`. +`-∫ x in Ioc b a, f x ∂μ` if `b ≤ a`. ## Implementation notes @@ -127,39 +40,9 @@ three possible intervals with the same endpoints for two reasons: [cumulative distribution function](https://en.wikipedia.org/wiki/Cumulative_distribution_function) of `μ`. -### `FTC_filter` class - -As explained above, many theorems in this file rely on the typeclass -`FTC_filter (a : ℝ) (l l' : filter ℝ)` to avoid code duplication. This typeclass combines four -assumptions: - -- `pure a ≤ l`; -- `l' ≤ 𝓝 a`; -- `l'` has a basis of measurable sets; -- if `u n` and `v n` tend to `l`, then for any `s ∈ l'`, `Ioc (u n) (v n)` is eventually included - in `s`. - -This typeclass has the following “real” instances: `(a, pure a, ⊥)`, `(a, 𝓝[≥] a, 𝓝[>] a)`, -`(a, 𝓝[≤] a, 𝓝[≤] a)`, `(a, 𝓝 a, 𝓝 a)`. -Furthermore, we have the following instances that are equal to the previously mentioned instances: -`(a, 𝓝[{a}] a, ⊥)` and `(a, 𝓝[univ] a, 𝓝[univ] a)`. -While the difference between `Ici a` and `Ioi a` doesn't matter for theorems about Lebesgue measure, -it becomes important in the versions of FTC about any locally finite measure if this measure has an -atom at one of the endpoints. - -### Combining one-sided and two-sided derivatives - -There are some `FTC_filter` instances where the fact that it is one-sided or -two-sided depends on the point, namely `(x, 𝓝[Icc a b] x, 𝓝[Icc a b] x)` -(resp. `(x, 𝓝[[a, b]] x, 𝓝[[a, b]] x)`, where `[a, b] = set.uIcc a b`), -with `x ∈ Icc a b` (resp. `x ∈ [a, b]`). -This results in a two-sided derivatives for `x ∈ Ioo a b` and one-sided derivatives for -`x ∈ {a, b}`. Other instances could be added when needed (in that case, one also needs to add -instances for `filter.is_measurably_generated` and `filter.tendsto_Ixx_class`). - ## Tags -integral, fundamental theorem of calculus, FTC-1, FTC-2, change of variables in integrals +integral -/ noncomputable theory @@ -171,7 +54,7 @@ open_locale classical topology filter ennreal big_operators interval nnreal variables {ι 𝕜 E F A : Type*} [normed_add_comm_group E] /-! -### Integrability at an interval +### Integrability on an interval -/ /-- A function `f` is called *interval integrable* with respect to a measure `μ` on an unordered @@ -1472,1343 +1355,5 @@ begin end end has_sum -/-! -### Fundamental theorem of calculus, part 1, for any measure - -In this section we prove a few lemmas that can be seen as versions of FTC-1 for interval integrals -w.r.t. any measure. Many theorems are formulated for one or two pairs of filters related by -`FTC_filter a l l'`. This typeclass has exactly four “real” instances: `(a, pure a, ⊥)`, -`(a, 𝓝[≥] a, 𝓝[>] a)`, `(a, 𝓝[≤] a, 𝓝[≤] a)`, `(a, 𝓝 a, 𝓝 a)`, and two instances -that are equal to the first and last “real” instances: `(a, 𝓝[{a}] a, ⊥)` and -`(a, 𝓝[univ] a, 𝓝[univ] a)`. We use this approach to avoid repeating arguments in many very similar -cases. Lean can automatically find both `a` and `l'` based on `l`. - -The most general theorem `measure_integral_sub_integral_sub_linear_is_o_of_tendsto_ae` can be seen -as a generalization of lemma `integral_has_strict_fderiv_at` below which states strict -differentiability of `∫ x in u..v, f x` in `(u, v)` at `(a, b)` for a measurable function `f` that -is integrable on `a..b` and is continuous at `a` and `b`. The lemma is generalized in three -directions: first, `measure_integral_sub_integral_sub_linear_is_o_of_tendsto_ae` deals with any -locally finite measure `μ`; second, it works for one-sided limits/derivatives; third, it assumes -only that `f` has finite limits almost surely at `a` and `b`. - -Namely, let `f` be a measurable function integrable on `a..b`. Let `(la, la')` be a pair of -`FTC_filter`s around `a`; let `(lb, lb')` be a pair of `FTC_filter`s around `b`. Suppose that `f` -has finite limits `ca` and `cb` at `la' ⊓ μ.ae` and `lb' ⊓ μ.ae`, respectively. Then -`∫ x in va..vb, f x ∂μ - ∫ x in ua..ub, f x ∂μ = ∫ x in ub..vb, cb ∂μ - ∫ x in ua..va, ca ∂μ + - o(‖∫ x in ua..va, (1:ℝ) ∂μ‖ + ‖∫ x in ub..vb, (1:ℝ) ∂μ‖)` -as `ua` and `va` tend to `la` while `ub` and `vb` tend to `lb`. - -This theorem is formulated with integral of constants instead of measures in the right hand sides -for two reasons: first, this way we avoid `min`/`max` in the statements; second, often it is -possible to write better `simp` lemmas for these integrals, see `integral_const` and -`integral_const_of_cdf`. - -In the next subsection we apply this theorem to prove various theorems about differentiability -of the integral w.r.t. Lebesgue measure. -/ - -/-- An auxiliary typeclass for the Fundamental theorem of calculus, part 1. It is used to formulate -theorems that work simultaneously for left and right one-sided derivatives of `∫ x in u..v, f x`. -/ -class FTC_filter (a : out_param ℝ) (outer : filter ℝ) (inner : out_param $ filter ℝ) - extends tendsto_Ixx_class Ioc outer inner : Prop := -(pure_le : pure a ≤ outer) -(le_nhds : inner ≤ 𝓝 a) -[meas_gen : is_measurably_generated inner] - -/- The `dangerous_instance` linter doesn't take `out_param`s into account, so it thinks that -`FTC_filter.to_tendsto_Ixx_class` is dangerous. Disable this linter using `nolint`. --/ -attribute [nolint dangerous_instance] FTC_filter.to_tendsto_Ixx_class - -namespace FTC_filter - -instance pure (a : ℝ) : FTC_filter a (pure a) ⊥ := -{ pure_le := le_rfl, - le_nhds := bot_le } - -instance nhds_within_singleton (a : ℝ) : FTC_filter a (𝓝[{a}] a) ⊥ := -by { rw [nhds_within, principal_singleton, inf_eq_right.2 (pure_le_nhds a)], apply_instance } - -lemma finite_at_inner {a : ℝ} (l : filter ℝ) {l'} [h : FTC_filter a l l'] - {μ : measure ℝ} [is_locally_finite_measure μ] : - μ.finite_at_filter l' := -(μ.finite_at_nhds a).filter_mono h.le_nhds - -instance nhds (a : ℝ) : FTC_filter a (𝓝 a) (𝓝 a) := -{ pure_le := pure_le_nhds a, - le_nhds := le_rfl } - -instance nhds_univ (a : ℝ) : FTC_filter a (𝓝[univ] a) (𝓝 a) := -by { rw nhds_within_univ, apply_instance } - -instance nhds_left (a : ℝ) : FTC_filter a (𝓝[≤] a) (𝓝[≤] a) := -{ pure_le := pure_le_nhds_within right_mem_Iic, - le_nhds := inf_le_left } - -instance nhds_right (a : ℝ) : FTC_filter a (𝓝[≥] a) (𝓝[>] a) := -{ pure_le := pure_le_nhds_within left_mem_Ici, - le_nhds := inf_le_left } - -instance nhds_Icc {x a b : ℝ} [h : fact (x ∈ Icc a b)] : - FTC_filter x (𝓝[Icc a b] x) (𝓝[Icc a b] x) := -{ pure_le := pure_le_nhds_within h.out, - le_nhds := inf_le_left } - -instance nhds_uIcc {x a b : ℝ} [h : fact (x ∈ [a, b])] : - FTC_filter x (𝓝[[a, b]] x) (𝓝[[a, b]] x) := -by { haveI : fact (x ∈ set.Icc (min a b) (max a b)) := h, exact FTC_filter.nhds_Icc } - -end FTC_filter - -open asymptotics - -section - -variables {f : ℝ → E} {a b : ℝ} {c ca cb : E} {l l' la la' lb lb' : filter ℝ} {lt : filter ι} - {μ : measure ℝ} {u v ua va ub vb : ι → ℝ} - -/-- Fundamental theorem of calculus-1, local version for any measure. -Let filters `l` and `l'` be related by `tendsto_Ixx_class Ioc`. -If `f` has a finite limit `c` at `l' ⊓ μ.ae`, where `μ` is a measure -finite at `l'`, then `∫ x in u..v, f x ∂μ = ∫ x in u..v, c ∂μ + o(∫ x in u..v, 1 ∂μ)` as both -`u` and `v` tend to `l`. - -See also `measure_integral_sub_linear_is_o_of_tendsto_ae` for a version assuming -`[FTC_filter a l l']` and `[is_locally_finite_measure μ]`. If `l` is one of `𝓝[≥] a`, -`𝓝[≤] a`, `𝓝 a`, then it's easier to apply the non-primed version. -The primed version also works, e.g., for `l = l' = at_top`. - -We use integrals of constants instead of measures because this way it is easier to formulate -a statement that works in both cases `u ≤ v` and `v ≤ u`. -/ -lemma measure_integral_sub_linear_is_o_of_tendsto_ae' - [is_measurably_generated l'] [tendsto_Ixx_class Ioc l l'] - (hfm : strongly_measurable_at_filter f l' μ) (hf : tendsto f (l' ⊓ μ.ae) (𝓝 c)) - (hl : μ.finite_at_filter l') - (hu : tendsto u lt l) (hv : tendsto v lt l) : - (λ t, ∫ x in u t..v t, f x ∂μ - ∫ x in u t..v t, c ∂μ) =o[lt] (λ t, ∫ x in u t..v t, (1:ℝ) ∂μ) := -begin - have A := hf.integral_sub_linear_is_o_ae hfm hl (hu.Ioc hv), - have B := hf.integral_sub_linear_is_o_ae hfm hl (hv.Ioc hu), - simp only [integral_const'], - convert (A.trans_le _).sub (B.trans_le _), - { ext t, - simp_rw [interval_integral, sub_smul], - abel }, - all_goals { intro t, cases le_total (u t) (v t) with huv huv; simp [huv] } -end - -/-- Fundamental theorem of calculus-1, local version for any measure. -Let filters `l` and `l'` be related by `tendsto_Ixx_class Ioc`. -If `f` has a finite limit `c` at `l ⊓ μ.ae`, where `μ` is a measure -finite at `l`, then `∫ x in u..v, f x ∂μ = μ (Ioc u v) • c + o(μ(Ioc u v))` as both -`u` and `v` tend to `l` so that `u ≤ v`. - -See also `measure_integral_sub_linear_is_o_of_tendsto_ae_of_le` for a version assuming -`[FTC_filter a l l']` and `[is_locally_finite_measure μ]`. If `l` is one of `𝓝[≥] a`, -`𝓝[≤] a`, `𝓝 a`, then it's easier to apply the non-primed version. -The primed version also works, e.g., for `l = l' = at_top`. -/ -lemma measure_integral_sub_linear_is_o_of_tendsto_ae_of_le' - [is_measurably_generated l'] [tendsto_Ixx_class Ioc l l'] - (hfm : strongly_measurable_at_filter f l' μ) (hf : tendsto f (l' ⊓ μ.ae) (𝓝 c)) - (hl : μ.finite_at_filter l') (hu : tendsto u lt l) (hv : tendsto v lt l) (huv : u ≤ᶠ[lt] v) : - (λ t, ∫ x in u t..v t, f x ∂μ - (μ (Ioc (u t) (v t))).to_real • c) =o[lt] - (λ t, (μ $ Ioc (u t) (v t)).to_real) := -(measure_integral_sub_linear_is_o_of_tendsto_ae' hfm hf hl hu hv).congr' - (huv.mono $ λ x hx, by simp [integral_const', hx]) - (huv.mono $ λ x hx, by simp [integral_const', hx]) - -/-- Fundamental theorem of calculus-1, local version for any measure. -Let filters `l` and `l'` be related by `tendsto_Ixx_class Ioc`. -If `f` has a finite limit `c` at `l ⊓ μ.ae`, where `μ` is a measure -finite at `l`, then `∫ x in u..v, f x ∂μ = -μ (Ioc v u) • c + o(μ(Ioc v u))` as both -`u` and `v` tend to `l` so that `v ≤ u`. - -See also `measure_integral_sub_linear_is_o_of_tendsto_ae_of_ge` for a version assuming -`[FTC_filter a l l']` and `[is_locally_finite_measure μ]`. If `l` is one of `𝓝[≥] a`, -`𝓝[≤] a`, `𝓝 a`, then it's easier to apply the non-primed version. -The primed version also works, e.g., for `l = l' = at_top`. -/ -lemma measure_integral_sub_linear_is_o_of_tendsto_ae_of_ge' - [is_measurably_generated l'] [tendsto_Ixx_class Ioc l l'] - (hfm : strongly_measurable_at_filter f l' μ) (hf : tendsto f (l' ⊓ μ.ae) (𝓝 c)) - (hl : μ.finite_at_filter l') (hu : tendsto u lt l) (hv : tendsto v lt l) (huv : v ≤ᶠ[lt] u) : - (λ t, ∫ x in u t..v t, f x ∂μ + (μ (Ioc (v t) (u t))).to_real • c) =o[lt] - (λ t, (μ $ Ioc (v t) (u t)).to_real) := -(measure_integral_sub_linear_is_o_of_tendsto_ae_of_le' hfm hf hl hv hu huv).neg_left.congr_left $ - λ t, by simp [integral_symm (u t), add_comm] - -section - -variables [is_locally_finite_measure μ] [FTC_filter a l l'] - -include a - -local attribute [instance] FTC_filter.meas_gen - -/-- Fundamental theorem of calculus-1, local version for any measure. -Let filters `l` and `l'` be related by `[FTC_filter a l l']`; let `μ` be a locally finite measure. -If `f` has a finite limit `c` at `l' ⊓ μ.ae`, then -`∫ x in u..v, f x ∂μ = ∫ x in u..v, c ∂μ + o(∫ x in u..v, 1 ∂μ)` as both `u` and `v` tend to `l`. - -See also `measure_integral_sub_linear_is_o_of_tendsto_ae'` for a version that also works, e.g., for -`l = l' = at_top`. - -We use integrals of constants instead of measures because this way it is easier to formulate -a statement that works in both cases `u ≤ v` and `v ≤ u`. -/ -lemma measure_integral_sub_linear_is_o_of_tendsto_ae (hfm : strongly_measurable_at_filter f l' μ) - (hf : tendsto f (l' ⊓ μ.ae) (𝓝 c)) (hu : tendsto u lt l) (hv : tendsto v lt l) : - (λ t, ∫ x in u t..v t, f x ∂μ - ∫ x in u t..v t, c ∂μ) =o[lt] (λ t, ∫ x in u t..v t, (1:ℝ) ∂μ) := -measure_integral_sub_linear_is_o_of_tendsto_ae' hfm hf (FTC_filter.finite_at_inner l) hu hv - -/-- Fundamental theorem of calculus-1, local version for any measure. -Let filters `l` and `l'` be related by `[FTC_filter a l l']`; let `μ` be a locally finite measure. -If `f` has a finite limit `c` at `l' ⊓ μ.ae`, then -`∫ x in u..v, f x ∂μ = μ (Ioc u v) • c + o(μ(Ioc u v))` as both `u` and `v` tend to `l`. - -See also `measure_integral_sub_linear_is_o_of_tendsto_ae_of_le'` for a version that also works, -e.g., for `l = l' = at_top`. -/ -lemma measure_integral_sub_linear_is_o_of_tendsto_ae_of_le - (hfm : strongly_measurable_at_filter f l' μ) (hf : tendsto f (l' ⊓ μ.ae) (𝓝 c)) - (hu : tendsto u lt l) (hv : tendsto v lt l) (huv : u ≤ᶠ[lt] v) : - (λ t, ∫ x in u t..v t, f x ∂μ - (μ (Ioc (u t) (v t))).to_real • c) =o[lt] - (λ t, (μ $ Ioc (u t) (v t)).to_real) := -measure_integral_sub_linear_is_o_of_tendsto_ae_of_le' hfm hf (FTC_filter.finite_at_inner l) - hu hv huv - -/-- Fundamental theorem of calculus-1, local version for any measure. -Let filters `l` and `l'` be related by `[FTC_filter a l l']`; let `μ` be a locally finite measure. -If `f` has a finite limit `c` at `l' ⊓ μ.ae`, then -`∫ x in u..v, f x ∂μ = -μ (Ioc v u) • c + o(μ(Ioc v u))` as both `u` and `v` tend to `l`. - -See also `measure_integral_sub_linear_is_o_of_tendsto_ae_of_ge'` for a version that also works, -e.g., for `l = l' = at_top`. -/ -lemma measure_integral_sub_linear_is_o_of_tendsto_ae_of_ge - (hfm : strongly_measurable_at_filter f l' μ) (hf : tendsto f (l' ⊓ μ.ae) (𝓝 c)) - (hu : tendsto u lt l) (hv : tendsto v lt l) (huv : v ≤ᶠ[lt] u) : - (λ t, ∫ x in u t..v t, f x ∂μ + (μ (Ioc (v t) (u t))).to_real • c) =o[lt] - (λ t, (μ $ Ioc (v t) (u t)).to_real) := -measure_integral_sub_linear_is_o_of_tendsto_ae_of_ge' hfm hf (FTC_filter.finite_at_inner l) - hu hv huv - -end - -local attribute [instance] FTC_filter.meas_gen - -variables [FTC_filter a la la'] [FTC_filter b lb lb'] [is_locally_finite_measure μ] - -/-- Fundamental theorem of calculus-1, strict derivative in both limits for a locally finite -measure. - -Let `f` be a measurable function integrable on `a..b`. Let `(la, la')` be a pair of `FTC_filter`s -around `a`; let `(lb, lb')` be a pair of `FTC_filter`s around `b`. Suppose that `f` has finite -limits `ca` and `cb` at `la' ⊓ μ.ae` and `lb' ⊓ μ.ae`, respectively. -Then `∫ x in va..vb, f x ∂μ - ∫ x in ua..ub, f x ∂μ = - ∫ x in ub..vb, cb ∂μ - ∫ x in ua..va, ca ∂μ + - o(‖∫ x in ua..va, (1:ℝ) ∂μ‖ + ‖∫ x in ub..vb, (1:ℝ) ∂μ‖)` -as `ua` and `va` tend to `la` while `ub` and `vb` tend to `lb`. --/ -lemma measure_integral_sub_integral_sub_linear_is_o_of_tendsto_ae - (hab : interval_integrable f μ a b) - (hmeas_a : strongly_measurable_at_filter f la' μ) - (hmeas_b : strongly_measurable_at_filter f lb' μ) - (ha_lim : tendsto f (la' ⊓ μ.ae) (𝓝 ca)) (hb_lim : tendsto f (lb' ⊓ μ.ae) (𝓝 cb)) - (hua : tendsto ua lt la) (hva : tendsto va lt la) - (hub : tendsto ub lt lb) (hvb : tendsto vb lt lb) : - (λ t, (∫ x in va t..vb t, f x ∂μ) - (∫ x in ua t..ub t, f x ∂μ) - - (∫ x in ub t..vb t, cb ∂μ - ∫ x in ua t..va t, ca ∂μ)) =o[lt] - (λ t, ‖∫ x in ua t..va t, (1:ℝ) ∂μ‖ + ‖∫ x in ub t..vb t, (1:ℝ) ∂μ‖) := -begin - refine - ((measure_integral_sub_linear_is_o_of_tendsto_ae hmeas_a ha_lim hua hva).neg_left.add_add - (measure_integral_sub_linear_is_o_of_tendsto_ae hmeas_b hb_lim hub hvb)).congr' - _ eventually_eq.rfl, - have A : ∀ᶠ t in lt, interval_integrable f μ (ua t) (va t) := - ha_lim.eventually_interval_integrable_ae hmeas_a (FTC_filter.finite_at_inner la) hua hva, - have A' : ∀ᶠ t in lt, interval_integrable f μ a (ua t) := - ha_lim.eventually_interval_integrable_ae hmeas_a (FTC_filter.finite_at_inner la) - (tendsto_const_pure.mono_right FTC_filter.pure_le) hua, - have B : ∀ᶠ t in lt, interval_integrable f μ (ub t) (vb t) := - hb_lim.eventually_interval_integrable_ae hmeas_b (FTC_filter.finite_at_inner lb) hub hvb, - have B' : ∀ᶠ t in lt, interval_integrable f μ b (ub t) := - hb_lim.eventually_interval_integrable_ae hmeas_b (FTC_filter.finite_at_inner lb) - (tendsto_const_pure.mono_right FTC_filter.pure_le) hub, - filter_upwards [A, A', B, B'] with _ ua_va a_ua ub_vb b_ub, - rw [← integral_interval_sub_interval_comm'], - { dsimp only [], abel, }, - exacts [ub_vb, ua_va, b_ub.symm.trans $ hab.symm.trans a_ua] -end - -/-- Fundamental theorem of calculus-1, strict derivative in right endpoint for a locally finite -measure. - -Let `f` be a measurable function integrable on `a..b`. Let `(lb, lb')` be a pair of `FTC_filter`s -around `b`. Suppose that `f` has a finite limit `c` at `lb' ⊓ μ.ae`. - -Then `∫ x in a..v, f x ∂μ - ∫ x in a..u, f x ∂μ = ∫ x in u..v, c ∂μ + o(∫ x in u..v, (1:ℝ) ∂μ)` -as `u` and `v` tend to `lb`. --/ -lemma measure_integral_sub_integral_sub_linear_is_o_of_tendsto_ae_right - (hab : interval_integrable f μ a b) (hmeas : strongly_measurable_at_filter f lb' μ) - (hf : tendsto f (lb' ⊓ μ.ae) (𝓝 c)) (hu : tendsto u lt lb) (hv : tendsto v lt lb) : - (λ t, ∫ x in a..v t, f x ∂μ - ∫ x in a..u t, f x ∂μ - ∫ x in u t..v t, c ∂μ) =o[lt] - (λ t, ∫ x in u t..v t, (1:ℝ) ∂μ) := -by simpa using measure_integral_sub_integral_sub_linear_is_o_of_tendsto_ae - hab strongly_measurable_at_bot hmeas ((tendsto_bot : tendsto _ ⊥ (𝓝 0)).mono_left inf_le_left) - hf (tendsto_const_pure : tendsto _ _ (pure a)) tendsto_const_pure hu hv - -/-- Fundamental theorem of calculus-1, strict derivative in left endpoint for a locally finite -measure. - -Let `f` be a measurable function integrable on `a..b`. Let `(la, la')` be a pair of `FTC_filter`s -around `a`. Suppose that `f` has a finite limit `c` at `la' ⊓ μ.ae`. - -Then `∫ x in v..b, f x ∂μ - ∫ x in u..b, f x ∂μ = -∫ x in u..v, c ∂μ + o(∫ x in u..v, (1:ℝ) ∂μ)` -as `u` and `v` tend to `la`. --/ -lemma measure_integral_sub_integral_sub_linear_is_o_of_tendsto_ae_left - (hab : interval_integrable f μ a b) (hmeas : strongly_measurable_at_filter f la' μ) - (hf : tendsto f (la' ⊓ μ.ae) (𝓝 c)) (hu : tendsto u lt la) (hv : tendsto v lt la) : - (λ t, ∫ x in v t..b, f x ∂μ - ∫ x in u t..b, f x ∂μ + ∫ x in u t..v t, c ∂μ) =o[lt] - (λ t, ∫ x in u t..v t, (1:ℝ) ∂μ) := -by simpa using measure_integral_sub_integral_sub_linear_is_o_of_tendsto_ae - hab hmeas strongly_measurable_at_bot hf ((tendsto_bot : tendsto _ ⊥ (𝓝 0)).mono_left inf_le_left) - hu hv (tendsto_const_pure : tendsto _ _ (pure b)) tendsto_const_pure - -end - -/-! -### Fundamental theorem of calculus-1 for Lebesgue measure - -In this section we restate theorems from the previous section for Lebesgue measure. -In particular, we prove that `∫ x in u..v, f x` is strictly differentiable in `(u, v)` -at `(a, b)` provided that `f` is integrable on `a..b` and is continuous at `a` and `b`. --/ - -variables {f : ℝ → E} {c ca cb : E} {l l' la la' lb lb' : filter ℝ} {lt : filter ι} - {a b z : ℝ} {u v ua ub va vb : ι → ℝ} [FTC_filter a la la'] [FTC_filter b lb lb'] - -/-! -#### Auxiliary `is_o` statements - -In this section we prove several lemmas that can be interpreted as strict differentiability of -`(u, v) ↦ ∫ x in u..v, f x ∂μ` in `u` and/or `v` at a filter. The statements use `is_o` because -we have no definition of `has_strict_(f)deriv_at_filter` in the library. --/ - -/-- Fundamental theorem of calculus-1, local version. If `f` has a finite limit `c` almost surely at -`l'`, where `(l, l')` is an `FTC_filter` pair around `a`, then -`∫ x in u..v, f x ∂μ = (v - u) • c + o (v - u)` as both `u` and `v` tend to `l`. -/ -lemma integral_sub_linear_is_o_of_tendsto_ae [FTC_filter a l l'] - (hfm : strongly_measurable_at_filter f l') (hf : tendsto f (l' ⊓ volume.ae) (𝓝 c)) - {u v : ι → ℝ} (hu : tendsto u lt l) (hv : tendsto v lt l) : - (λ t, (∫ x in u t..v t, f x) - (v t - u t) • c) =o[lt] (v - u) := -by simpa [integral_const] using measure_integral_sub_linear_is_o_of_tendsto_ae hfm hf hu hv - -/-- Fundamental theorem of calculus-1, strict differentiability at filter in both endpoints. -If `f` is a measurable function integrable on `a..b`, `(la, la')` is an `FTC_filter` pair around -`a`, and `(lb, lb')` is an `FTC_filter` pair around `b`, and `f` has finite limits `ca` and `cb` -almost surely at `la'` and `lb'`, respectively, then -`(∫ x in va..vb, f x) - ∫ x in ua..ub, f x = (vb - ub) • cb - (va - ua) • ca + - o(‖va - ua‖ + ‖vb - ub‖)` as `ua` and `va` tend to `la` while `ub` and `vb` tend to `lb`. - -This lemma could've been formulated using `has_strict_fderiv_at_filter` if we had this -definition. -/ -lemma integral_sub_integral_sub_linear_is_o_of_tendsto_ae - (hab : interval_integrable f volume a b) - (hmeas_a : strongly_measurable_at_filter f la') (hmeas_b : strongly_measurable_at_filter f lb') - (ha_lim : tendsto f (la' ⊓ volume.ae) (𝓝 ca)) (hb_lim : tendsto f (lb' ⊓ volume.ae) (𝓝 cb)) - (hua : tendsto ua lt la) (hva : tendsto va lt la) - (hub : tendsto ub lt lb) (hvb : tendsto vb lt lb) : - (λ t, (∫ x in va t..vb t, f x) - (∫ x in ua t..ub t, f x) - - ((vb t - ub t) • cb - (va t - ua t) • ca)) =o[lt] (λ t, ‖va t - ua t‖ + ‖vb t - ub t‖) := -by simpa [integral_const] - using measure_integral_sub_integral_sub_linear_is_o_of_tendsto_ae hab hmeas_a hmeas_b - ha_lim hb_lim hua hva hub hvb - -/-- Fundamental theorem of calculus-1, strict differentiability at filter in both endpoints. -If `f` is a measurable function integrable on `a..b`, `(lb, lb')` is an `FTC_filter` pair -around `b`, and `f` has a finite limit `c` almost surely at `lb'`, then -`(∫ x in a..v, f x) - ∫ x in a..u, f x = (v - u) • c + o(‖v - u‖)` as `u` and `v` tend to `lb`. - -This lemma could've been formulated using `has_strict_deriv_at_filter` if we had this definition. -/ -lemma integral_sub_integral_sub_linear_is_o_of_tendsto_ae_right - (hab : interval_integrable f volume a b) (hmeas : strongly_measurable_at_filter f lb') - (hf : tendsto f (lb' ⊓ volume.ae) (𝓝 c)) (hu : tendsto u lt lb) (hv : tendsto v lt lb) : - (λ t, (∫ x in a..v t, f x) - (∫ x in a..u t, f x) - (v t - u t) • c) =o[lt] (v - u) := -by simpa only [integral_const, smul_eq_mul, mul_one] using - measure_integral_sub_integral_sub_linear_is_o_of_tendsto_ae_right hab hmeas hf hu hv - -/-- Fundamental theorem of calculus-1, strict differentiability at filter in both endpoints. -If `f` is a measurable function integrable on `a..b`, `(la, la')` is an `FTC_filter` pair -around `a`, and `f` has a finite limit `c` almost surely at `la'`, then -`(∫ x in v..b, f x) - ∫ x in u..b, f x = -(v - u) • c + o(‖v - u‖)` as `u` and `v` tend to `la`. - -This lemma could've been formulated using `has_strict_deriv_at_filter` if we had this definition. -/ -lemma integral_sub_integral_sub_linear_is_o_of_tendsto_ae_left - (hab : interval_integrable f volume a b) (hmeas : strongly_measurable_at_filter f la') - (hf : tendsto f (la' ⊓ volume.ae) (𝓝 c)) (hu : tendsto u lt la) (hv : tendsto v lt la) : - (λ t, (∫ x in v t..b, f x) - (∫ x in u t..b, f x) + (v t - u t) • c) =o[lt] (v - u) := -by simpa only [integral_const, smul_eq_mul, mul_one] using - measure_integral_sub_integral_sub_linear_is_o_of_tendsto_ae_left hab hmeas hf hu hv - -open continuous_linear_map (fst snd smul_right sub_apply smul_right_apply coe_fst' coe_snd' map_sub) - -/-! -#### Strict differentiability - -In this section we prove that for a measurable function `f` integrable on `a..b`, - -* `integral_has_strict_fderiv_at_of_tendsto_ae`: the function `(u, v) ↦ ∫ x in u..v, f x` has - derivative `(u, v) ↦ v • cb - u • ca` at `(a, b)` in the sense of strict differentiability - provided that `f` tends to `ca` and `cb` almost surely as `x` tendsto to `a` and `b`, - respectively; - -* `integral_has_strict_fderiv_at`: the function `(u, v) ↦ ∫ x in u..v, f x` has - derivative `(u, v) ↦ v • f b - u • f a` at `(a, b)` in the sense of strict differentiability - provided that `f` is continuous at `a` and `b`; - -* `integral_has_strict_deriv_at_of_tendsto_ae_right`: the function `u ↦ ∫ x in a..u, f x` has - derivative `c` at `b` in the sense of strict differentiability provided that `f` tends to `c` - almost surely as `x` tends to `b`; - -* `integral_has_strict_deriv_at_right`: the function `u ↦ ∫ x in a..u, f x` has derivative `f b` at - `b` in the sense of strict differentiability provided that `f` is continuous at `b`; - -* `integral_has_strict_deriv_at_of_tendsto_ae_left`: the function `u ↦ ∫ x in u..b, f x` has - derivative `-c` at `a` in the sense of strict differentiability provided that `f` tends to `c` - almost surely as `x` tends to `a`; - -* `integral_has_strict_deriv_at_left`: the function `u ↦ ∫ x in u..b, f x` has derivative `-f a` at - `a` in the sense of strict differentiability provided that `f` is continuous at `a`. --/ - -/-- Fundamental theorem of calculus-1: if `f : ℝ → E` is integrable on `a..b` and `f x` has finite -limits `ca` and `cb` almost surely as `x` tends to `a` and `b`, respectively, then -`(u, v) ↦ ∫ x in u..v, f x` has derivative `(u, v) ↦ v • cb - u • ca` at `(a, b)` -in the sense of strict differentiability. -/ -lemma integral_has_strict_fderiv_at_of_tendsto_ae - (hf : interval_integrable f volume a b) - (hmeas_a : strongly_measurable_at_filter f (𝓝 a)) - (hmeas_b : strongly_measurable_at_filter f (𝓝 b)) - (ha : tendsto f (𝓝 a ⊓ volume.ae) (𝓝 ca)) (hb : tendsto f (𝓝 b ⊓ volume.ae) (𝓝 cb)) : - has_strict_fderiv_at (λ p : ℝ × ℝ, ∫ x in p.1..p.2, f x) - ((snd ℝ ℝ ℝ).smul_right cb - (fst ℝ ℝ ℝ).smul_right ca) (a, b) := -begin - have := integral_sub_integral_sub_linear_is_o_of_tendsto_ae hf hmeas_a hmeas_b ha hb - ((continuous_fst.comp continuous_snd).tendsto ((a, b), (a, b))) - ((continuous_fst.comp continuous_fst).tendsto ((a, b), (a, b))) - ((continuous_snd.comp continuous_snd).tendsto ((a, b), (a, b))) - ((continuous_snd.comp continuous_fst).tendsto ((a, b), (a, b))), - refine (this.congr_left _).trans_is_O _, - { intro x, simp [sub_smul] }, - { exact is_O_fst_prod.norm_left.add is_O_snd_prod.norm_left } -end - -/-- Fundamental theorem of calculus-1: if `f : ℝ → E` is integrable on `a..b` and `f` is continuous -at `a` and `b`, then `(u, v) ↦ ∫ x in u..v, f x` has derivative `(u, v) ↦ v • cb - u • ca` -at `(a, b)` in the sense of strict differentiability. -/ -lemma integral_has_strict_fderiv_at - (hf : interval_integrable f volume a b) - (hmeas_a : strongly_measurable_at_filter f (𝓝 a)) - (hmeas_b : strongly_measurable_at_filter f (𝓝 b)) - (ha : continuous_at f a) (hb : continuous_at f b) : - has_strict_fderiv_at (λ p : ℝ × ℝ, ∫ x in p.1..p.2, f x) - ((snd ℝ ℝ ℝ).smul_right (f b) - (fst ℝ ℝ ℝ).smul_right (f a)) (a, b) := -integral_has_strict_fderiv_at_of_tendsto_ae hf hmeas_a hmeas_b - (ha.mono_left inf_le_left) (hb.mono_left inf_le_left) - -/-- **First Fundamental Theorem of Calculus**: if `f : ℝ → E` is integrable on `a..b` and `f x` has -a finite limit `c` almost surely at `b`, then `u ↦ ∫ x in a..u, f x` has derivative `c` at `b` in -the sense of strict differentiability. -/ -lemma integral_has_strict_deriv_at_of_tendsto_ae_right - (hf : interval_integrable f volume a b) (hmeas : strongly_measurable_at_filter f (𝓝 b)) - (hb : tendsto f (𝓝 b ⊓ volume.ae) (𝓝 c)) : has_strict_deriv_at (λ u, ∫ x in a..u, f x) c b := -integral_sub_integral_sub_linear_is_o_of_tendsto_ae_right hf hmeas hb continuous_at_snd - continuous_at_fst - -/-- Fundamental theorem of calculus-1: if `f : ℝ → E` is integrable on `a..b` and `f` is continuous -at `b`, then `u ↦ ∫ x in a..u, f x` has derivative `f b` at `b` in the sense of strict -differentiability. -/ -lemma integral_has_strict_deriv_at_right - (hf : interval_integrable f volume a b) (hmeas : strongly_measurable_at_filter f (𝓝 b)) - (hb : continuous_at f b) : has_strict_deriv_at (λ u, ∫ x in a..u, f x) (f b) b := -integral_has_strict_deriv_at_of_tendsto_ae_right hf hmeas (hb.mono_left inf_le_left) - -/-- Fundamental theorem of calculus-1: if `f : ℝ → E` is integrable on `a..b` and `f x` has a finite -limit `c` almost surely at `a`, then `u ↦ ∫ x in u..b, f x` has derivative `-c` at `a` in the sense -of strict differentiability. -/ -lemma integral_has_strict_deriv_at_of_tendsto_ae_left - (hf : interval_integrable f volume a b) (hmeas : strongly_measurable_at_filter f (𝓝 a)) - (ha : tendsto f (𝓝 a ⊓ volume.ae) (𝓝 c)) : has_strict_deriv_at (λ u, ∫ x in u..b, f x) (-c) a := -by simpa only [← integral_symm] - using (integral_has_strict_deriv_at_of_tendsto_ae_right hf.symm hmeas ha).neg - -/-- Fundamental theorem of calculus-1: if `f : ℝ → E` is integrable on `a..b` and `f` is continuous -at `a`, then `u ↦ ∫ x in u..b, f x` has derivative `-f a` at `a` in the sense of strict -differentiability. -/ -lemma integral_has_strict_deriv_at_left - (hf : interval_integrable f volume a b) (hmeas : strongly_measurable_at_filter f (𝓝 a)) - (ha : continuous_at f a) : has_strict_deriv_at (λ u, ∫ x in u..b, f x) (-f a) a := -by simpa only [← integral_symm] using (integral_has_strict_deriv_at_right hf.symm hmeas ha).neg - -/-- Fundamental theorem of calculus-1: if `f : ℝ → E` is continuous, then `u ↦ ∫ x in a..u, f x` -has derivative `f b` at `b` in the sense of strict differentiability. -/ -lemma _root_.continuous.integral_has_strict_deriv_at {f : ℝ → E} (hf : continuous f) (a b : ℝ) : - has_strict_deriv_at (λ u, ∫ (x : ℝ) in a..u, f x) (f b) b := -integral_has_strict_deriv_at_right (hf.interval_integrable _ _) - (hf.strongly_measurable_at_filter _ _) hf.continuous_at - -/-- Fundamental theorem of calculus-1: if `f : ℝ → E` is continuous, then the derivative -of `u ↦ ∫ x in a..u, f x` at `b` is `f b`. -/ -lemma _root_.continuous.deriv_integral (f : ℝ → E) (hf : continuous f) (a b : ℝ) : - deriv (λ u, ∫ (x : ℝ) in a..u, f x) b = f b := -(hf.integral_has_strict_deriv_at a b).has_deriv_at.deriv - -/-! -#### Fréchet differentiability - -In this subsection we restate results from the previous subsection in terms of `has_fderiv_at`, -`has_deriv_at`, `fderiv`, and `deriv`. --/ - -/-- Fundamental theorem of calculus-1: if `f : ℝ → E` is integrable on `a..b` and `f x` has finite -limits `ca` and `cb` almost surely as `x` tends to `a` and `b`, respectively, then -`(u, v) ↦ ∫ x in u..v, f x` has derivative `(u, v) ↦ v • cb - u • ca` at `(a, b)`. -/ -lemma integral_has_fderiv_at_of_tendsto_ae (hf : interval_integrable f volume a b) - (hmeas_a : strongly_measurable_at_filter f (𝓝 a)) - (hmeas_b : strongly_measurable_at_filter f (𝓝 b)) - (ha : tendsto f (𝓝 a ⊓ volume.ae) (𝓝 ca)) (hb : tendsto f (𝓝 b ⊓ volume.ae) (𝓝 cb)) : - has_fderiv_at (λ p : ℝ × ℝ, ∫ x in p.1..p.2, f x) - ((snd ℝ ℝ ℝ).smul_right cb - (fst ℝ ℝ ℝ).smul_right ca) (a, b) := -(integral_has_strict_fderiv_at_of_tendsto_ae hf hmeas_a hmeas_b ha hb).has_fderiv_at - -/-- Fundamental theorem of calculus-1: if `f : ℝ → E` is integrable on `a..b` and `f` is continuous -at `a` and `b`, then `(u, v) ↦ ∫ x in u..v, f x` has derivative `(u, v) ↦ v • cb - u • ca` -at `(a, b)`. -/ -lemma integral_has_fderiv_at (hf : interval_integrable f volume a b) - (hmeas_a : strongly_measurable_at_filter f (𝓝 a)) - (hmeas_b : strongly_measurable_at_filter f (𝓝 b)) - (ha : continuous_at f a) (hb : continuous_at f b) : - has_fderiv_at (λ p : ℝ × ℝ, ∫ x in p.1..p.2, f x) - ((snd ℝ ℝ ℝ).smul_right (f b) - (fst ℝ ℝ ℝ).smul_right (f a)) (a, b) := -(integral_has_strict_fderiv_at hf hmeas_a hmeas_b ha hb).has_fderiv_at - -/-- Fundamental theorem of calculus-1: if `f : ℝ → E` is integrable on `a..b` and `f x` has finite -limits `ca` and `cb` almost surely as `x` tends to `a` and `b`, respectively, then `fderiv` -derivative of `(u, v) ↦ ∫ x in u..v, f x` at `(a, b)` equals `(u, v) ↦ v • cb - u • ca`. -/ -lemma fderiv_integral_of_tendsto_ae (hf : interval_integrable f volume a b) - (hmeas_a : strongly_measurable_at_filter f (𝓝 a)) - (hmeas_b : strongly_measurable_at_filter f (𝓝 b)) - (ha : tendsto f (𝓝 a ⊓ volume.ae) (𝓝 ca)) (hb : tendsto f (𝓝 b ⊓ volume.ae) (𝓝 cb)) : - fderiv ℝ (λ p : ℝ × ℝ, ∫ x in p.1..p.2, f x) (a, b) = - (snd ℝ ℝ ℝ).smul_right cb - (fst ℝ ℝ ℝ).smul_right ca := -(integral_has_fderiv_at_of_tendsto_ae hf hmeas_a hmeas_b ha hb).fderiv - -/-- Fundamental theorem of calculus-1: if `f : ℝ → E` is integrable on `a..b` and `f` is continuous -at `a` and `b`, then `fderiv` derivative of `(u, v) ↦ ∫ x in u..v, f x` at `(a, b)` equals `(u, v) ↦ -v • cb - u • ca`. -/ -lemma fderiv_integral (hf : interval_integrable f volume a b) - (hmeas_a : strongly_measurable_at_filter f (𝓝 a)) - (hmeas_b : strongly_measurable_at_filter f (𝓝 b)) - (ha : continuous_at f a) (hb : continuous_at f b) : - fderiv ℝ (λ p : ℝ × ℝ, ∫ x in p.1..p.2, f x) (a, b) = - (snd ℝ ℝ ℝ).smul_right (f b) - (fst ℝ ℝ ℝ).smul_right (f a) := -(integral_has_fderiv_at hf hmeas_a hmeas_b ha hb).fderiv - -/-- Fundamental theorem of calculus-1: if `f : ℝ → E` is integrable on `a..b` and `f x` has a finite -limit `c` almost surely at `b`, then `u ↦ ∫ x in a..u, f x` has derivative `c` at `b`. -/ -lemma integral_has_deriv_at_of_tendsto_ae_right - (hf : interval_integrable f volume a b) (hmeas : strongly_measurable_at_filter f (𝓝 b)) - (hb : tendsto f (𝓝 b ⊓ volume.ae) (𝓝 c)) : has_deriv_at (λ u, ∫ x in a..u, f x) c b := -(integral_has_strict_deriv_at_of_tendsto_ae_right hf hmeas hb).has_deriv_at - -/-- Fundamental theorem of calculus-1: if `f : ℝ → E` is integrable on `a..b` and `f` is continuous -at `b`, then `u ↦ ∫ x in a..u, f x` has derivative `f b` at `b`. -/ -lemma integral_has_deriv_at_right - (hf : interval_integrable f volume a b) (hmeas : strongly_measurable_at_filter f (𝓝 b)) - (hb : continuous_at f b) : has_deriv_at (λ u, ∫ x in a..u, f x) (f b) b := -(integral_has_strict_deriv_at_right hf hmeas hb).has_deriv_at - -/-- Fundamental theorem of calculus: if `f : ℝ → E` is integrable on `a..b` and `f` has a finite -limit `c` almost surely at `b`, then the derivative of `u ↦ ∫ x in a..u, f x` at `b` equals `c`. -/ -lemma deriv_integral_of_tendsto_ae_right - (hf : interval_integrable f volume a b) (hmeas : strongly_measurable_at_filter f (𝓝 b)) - (hb : tendsto f (𝓝 b ⊓ volume.ae) (𝓝 c)) : deriv (λ u, ∫ x in a..u, f x) b = c := -(integral_has_deriv_at_of_tendsto_ae_right hf hmeas hb).deriv - -/-- Fundamental theorem of calculus: if `f : ℝ → E` is integrable on `a..b` and `f` is continuous -at `b`, then the derivative of `u ↦ ∫ x in a..u, f x` at `b` equals `f b`. -/ -lemma deriv_integral_right - (hf : interval_integrable f volume a b) (hmeas : strongly_measurable_at_filter f (𝓝 b)) - (hb : continuous_at f b) : - deriv (λ u, ∫ x in a..u, f x) b = f b := -(integral_has_deriv_at_right hf hmeas hb).deriv - -/-- Fundamental theorem of calculus-1: if `f : ℝ → E` is integrable on `a..b` and `f x` has a finite -limit `c` almost surely at `a`, then `u ↦ ∫ x in u..b, f x` has derivative `-c` at `a`. -/ -lemma integral_has_deriv_at_of_tendsto_ae_left - (hf : interval_integrable f volume a b) (hmeas : strongly_measurable_at_filter f (𝓝 a)) - (ha : tendsto f (𝓝 a ⊓ volume.ae) (𝓝 c)) : has_deriv_at (λ u, ∫ x in u..b, f x) (-c) a := -(integral_has_strict_deriv_at_of_tendsto_ae_left hf hmeas ha).has_deriv_at - -/-- Fundamental theorem of calculus-1: if `f : ℝ → E` is integrable on `a..b` and `f` is continuous -at `a`, then `u ↦ ∫ x in u..b, f x` has derivative `-f a` at `a`. -/ -lemma integral_has_deriv_at_left - (hf : interval_integrable f volume a b) (hmeas : strongly_measurable_at_filter f (𝓝 a)) - (ha : continuous_at f a) : - has_deriv_at (λ u, ∫ x in u..b, f x) (-f a) a := -(integral_has_strict_deriv_at_left hf hmeas ha).has_deriv_at - -/-- Fundamental theorem of calculus: if `f : ℝ → E` is integrable on `a..b` and `f` has a finite -limit `c` almost surely at `a`, then the derivative of `u ↦ ∫ x in u..b, f x` at `a` equals `-c`. -/ -lemma deriv_integral_of_tendsto_ae_left - (hf : interval_integrable f volume a b) (hmeas : strongly_measurable_at_filter f (𝓝 a)) - (hb : tendsto f (𝓝 a ⊓ volume.ae) (𝓝 c)) : deriv (λ u, ∫ x in u..b, f x) a = -c := -(integral_has_deriv_at_of_tendsto_ae_left hf hmeas hb).deriv - -/-- Fundamental theorem of calculus: if `f : ℝ → E` is integrable on `a..b` and `f` is continuous -at `a`, then the derivative of `u ↦ ∫ x in u..b, f x` at `a` equals `-f a`. -/ -lemma deriv_integral_left - (hf : interval_integrable f volume a b) (hmeas : strongly_measurable_at_filter f (𝓝 a)) - (hb : continuous_at f a) : - deriv (λ u, ∫ x in u..b, f x) a = -f a := -(integral_has_deriv_at_left hf hmeas hb).deriv - -/-! -#### One-sided derivatives --/ - -/-- Let `f` be a measurable function integrable on `a..b`. The function `(u, v) ↦ ∫ x in u..v, f x` -has derivative `(u, v) ↦ v • cb - u • ca` within `s × t` at `(a, b)`, where -`s ∈ {Iic a, {a}, Ici a, univ}` and `t ∈ {Iic b, {b}, Ici b, univ}` provided that `f` tends to `ca` -and `cb` almost surely at the filters `la` and `lb` from the following table. - -| `s` | `la` | `t` | `lb` | -| ------- | ---- | --- | ---- | -| `Iic a` | `𝓝[≤] a` | `Iic b` | `𝓝[≤] b` | -| `Ici a` | `𝓝[>] a` | `Ici b` | `𝓝[>] b` | -| `{a}` | `⊥` | `{b}` | `⊥` | -| `univ` | `𝓝 a` | `univ` | `𝓝 b` | --/ -lemma integral_has_fderiv_within_at_of_tendsto_ae - (hf : interval_integrable f volume a b) - {s t : set ℝ} [FTC_filter a (𝓝[s] a) la] [FTC_filter b (𝓝[t] b) lb] - (hmeas_a : strongly_measurable_at_filter f la) (hmeas_b : strongly_measurable_at_filter f lb) - (ha : tendsto f (la ⊓ volume.ae) (𝓝 ca)) (hb : tendsto f (lb ⊓ volume.ae) (𝓝 cb)) : - has_fderiv_within_at (λ p : ℝ × ℝ, ∫ x in p.1..p.2, f x) - ((snd ℝ ℝ ℝ).smul_right cb - (fst ℝ ℝ ℝ).smul_right ca) (s ×ˢ t) (a, b) := -begin - rw [has_fderiv_within_at, nhds_within_prod_eq], - have := integral_sub_integral_sub_linear_is_o_of_tendsto_ae hf hmeas_a hmeas_b ha hb - (tendsto_const_pure.mono_right FTC_filter.pure_le : tendsto _ _ (𝓝[s] a)) tendsto_fst - (tendsto_const_pure.mono_right FTC_filter.pure_le : tendsto _ _ (𝓝[t] b)) tendsto_snd, - refine (this.congr_left _).trans_is_O _, - { intro x, simp [sub_smul] }, - { exact is_O_fst_prod.norm_left.add is_O_snd_prod.norm_left } -end - -/-- Let `f` be a measurable function integrable on `a..b`. The function `(u, v) ↦ ∫ x in u..v, f x` -has derivative `(u, v) ↦ v • f b - u • f a` within `s × t` at `(a, b)`, where -`s ∈ {Iic a, {a}, Ici a, univ}` and `t ∈ {Iic b, {b}, Ici b, univ}` provided that `f` tends to -`f a` and `f b` at the filters `la` and `lb` from the following table. In most cases this assumption -is definitionally equal `continuous_at f _` or `continuous_within_at f _ _`. - -| `s` | `la` | `t` | `lb` | -| ------- | ---- | --- | ---- | -| `Iic a` | `𝓝[≤] a` | `Iic b` | `𝓝[≤] b` | -| `Ici a` | `𝓝[>] a` | `Ici b` | `𝓝[>] b` | -| `{a}` | `⊥` | `{b}` | `⊥` | -| `univ` | `𝓝 a` | `univ` | `𝓝 b` | --/ -lemma integral_has_fderiv_within_at - (hf : interval_integrable f volume a b) - (hmeas_a : strongly_measurable_at_filter f la) (hmeas_b : strongly_measurable_at_filter f lb) - {s t : set ℝ} [FTC_filter a (𝓝[s] a) la] [FTC_filter b (𝓝[t] b) lb] - (ha : tendsto f la (𝓝 $ f a)) (hb : tendsto f lb (𝓝 $ f b)) : - has_fderiv_within_at (λ p : ℝ × ℝ, ∫ x in p.1..p.2, f x) - ((snd ℝ ℝ ℝ).smul_right (f b) - (fst ℝ ℝ ℝ).smul_right (f a)) (s ×ˢ t) (a, b) := -integral_has_fderiv_within_at_of_tendsto_ae hf hmeas_a hmeas_b (ha.mono_left inf_le_left) - (hb.mono_left inf_le_left) - -/-- An auxiliary tactic closing goals `unique_diff_within_at ℝ s a` where -`s ∈ {Iic a, Ici a, univ}`. -/ -meta def unique_diff_within_at_Ici_Iic_univ : tactic unit := -`[apply_rules [unique_diff_on.unique_diff_within_at, unique_diff_on_Ici, unique_diff_on_Iic, - left_mem_Ici, right_mem_Iic, unique_diff_within_at_univ]] - -/-- Let `f` be a measurable function integrable on `a..b`. Choose `s ∈ {Iic a, Ici a, univ}` -and `t ∈ {Iic b, Ici b, univ}`. Suppose that `f` tends to `ca` and `cb` almost surely at the filters -`la` and `lb` from the table below. Then `fderiv_within ℝ (λ p, ∫ x in p.1..p.2, f x) (s ×ˢ t)` -is equal to `(u, v) ↦ u • cb - v • ca`. - -| `s` | `la` | `t` | `lb` | -| ------- | ---- | --- | ---- | -| `Iic a` | `𝓝[≤] a` | `Iic b` | `𝓝[≤] b` | -| `Ici a` | `𝓝[>] a` | `Ici b` | `𝓝[>] b` | -| `{a}` | `⊥` | `{b}` | `⊥` | -| `univ` | `𝓝 a` | `univ` | `𝓝 b` | --/ -lemma fderiv_within_integral_of_tendsto_ae - (hf : interval_integrable f volume a b) - (hmeas_a : strongly_measurable_at_filter f la) (hmeas_b : strongly_measurable_at_filter f lb) - {s t : set ℝ} [FTC_filter a (𝓝[s] a) la] [FTC_filter b (𝓝[t] b) lb] - (ha : tendsto f (la ⊓ volume.ae) (𝓝 ca)) (hb : tendsto f (lb ⊓ volume.ae) (𝓝 cb)) - (hs : unique_diff_within_at ℝ s a . unique_diff_within_at_Ici_Iic_univ) - (ht : unique_diff_within_at ℝ t b . unique_diff_within_at_Ici_Iic_univ) : - fderiv_within ℝ (λ p : ℝ × ℝ, ∫ x in p.1..p.2, f x) (s ×ˢ t) (a, b) = - ((snd ℝ ℝ ℝ).smul_right cb - (fst ℝ ℝ ℝ).smul_right ca) := -(integral_has_fderiv_within_at_of_tendsto_ae hf hmeas_a hmeas_b ha hb).fderiv_within $ hs.prod ht - -/-- Fundamental theorem of calculus: if `f : ℝ → E` is integrable on `a..b` and `f x` has a finite -limit `c` almost surely as `x` tends to `b` from the right or from the left, -then `u ↦ ∫ x in a..u, f x` has right (resp., left) derivative `c` at `b`. -/ -lemma integral_has_deriv_within_at_of_tendsto_ae_right - (hf : interval_integrable f volume a b) {s t : set ℝ} [FTC_filter b (𝓝[s] b) (𝓝[t] b)] - (hmeas : strongly_measurable_at_filter f (𝓝[t] b)) (hb : tendsto f (𝓝[t] b ⊓ volume.ae) (𝓝 c)) : - has_deriv_within_at (λ u, ∫ x in a..u, f x) c s b := -integral_sub_integral_sub_linear_is_o_of_tendsto_ae_right hf hmeas hb - (tendsto_const_pure.mono_right FTC_filter.pure_le) tendsto_id - -/-- Fundamental theorem of calculus: if `f : ℝ → E` is integrable on `a..b` and `f x` is continuous -from the left or from the right at `b`, then `u ↦ ∫ x in a..u, f x` has left (resp., right) -derivative `f b` at `b`. -/ -lemma integral_has_deriv_within_at_right - (hf : interval_integrable f volume a b) {s t : set ℝ} [FTC_filter b (𝓝[s] b) (𝓝[t] b)] - (hmeas : strongly_measurable_at_filter f (𝓝[t] b)) (hb : continuous_within_at f t b) : - has_deriv_within_at (λ u, ∫ x in a..u, f x) (f b) s b := -integral_has_deriv_within_at_of_tendsto_ae_right hf hmeas (hb.mono_left inf_le_left) - -/-- Fundamental theorem of calculus: if `f : ℝ → E` is integrable on `a..b` and `f x` has a finite -limit `c` almost surely as `x` tends to `b` from the right or from the left, then the right -(resp., left) derivative of `u ↦ ∫ x in a..u, f x` at `b` equals `c`. -/ -lemma deriv_within_integral_of_tendsto_ae_right - (hf : interval_integrable f volume a b) {s t : set ℝ} [FTC_filter b (𝓝[s] b) (𝓝[t] b)] - (hmeas: strongly_measurable_at_filter f (𝓝[t] b)) (hb : tendsto f (𝓝[t] b ⊓ volume.ae) (𝓝 c)) - (hs : unique_diff_within_at ℝ s b . unique_diff_within_at_Ici_Iic_univ) : - deriv_within (λ u, ∫ x in a..u, f x) s b = c := -(integral_has_deriv_within_at_of_tendsto_ae_right hf hmeas hb).deriv_within hs - -/-- Fundamental theorem of calculus: if `f : ℝ → E` is integrable on `a..b` and `f x` is continuous -on the right or on the left at `b`, then the right (resp., left) derivative of -`u ↦ ∫ x in a..u, f x` at `b` equals `f b`. -/ -lemma deriv_within_integral_right - (hf : interval_integrable f volume a b) {s t : set ℝ} [FTC_filter b (𝓝[s] b) (𝓝[t] b)] - (hmeas : strongly_measurable_at_filter f (𝓝[t] b)) (hb : continuous_within_at f t b) - (hs : unique_diff_within_at ℝ s b . unique_diff_within_at_Ici_Iic_univ) : - deriv_within (λ u, ∫ x in a..u, f x) s b = f b := -(integral_has_deriv_within_at_right hf hmeas hb).deriv_within hs - -/-- Fundamental theorem of calculus: if `f : ℝ → E` is integrable on `a..b` and `f x` has a finite -limit `c` almost surely as `x` tends to `a` from the right or from the left, -then `u ↦ ∫ x in u..b, f x` has right (resp., left) derivative `-c` at `a`. -/ -lemma integral_has_deriv_within_at_of_tendsto_ae_left - (hf : interval_integrable f volume a b) {s t : set ℝ} [FTC_filter a (𝓝[s] a) (𝓝[t] a)] - (hmeas : strongly_measurable_at_filter f (𝓝[t] a)) - (ha : tendsto f (𝓝[t] a ⊓ volume.ae) (𝓝 c)) : - has_deriv_within_at (λ u, ∫ x in u..b, f x) (-c) s a := -by { simp only [integral_symm b], - exact (integral_has_deriv_within_at_of_tendsto_ae_right hf.symm hmeas ha).neg } - -/-- Fundamental theorem of calculus: if `f : ℝ → E` is integrable on `a..b` and `f x` is continuous -from the left or from the right at `a`, then `u ↦ ∫ x in u..b, f x` has left (resp., right) -derivative `-f a` at `a`. -/ -lemma integral_has_deriv_within_at_left - (hf : interval_integrable f volume a b) {s t : set ℝ} [FTC_filter a (𝓝[s] a) (𝓝[t] a)] - (hmeas : strongly_measurable_at_filter f (𝓝[t] a)) (ha : continuous_within_at f t a) : - has_deriv_within_at (λ u, ∫ x in u..b, f x) (-f a) s a := -integral_has_deriv_within_at_of_tendsto_ae_left hf hmeas (ha.mono_left inf_le_left) - -/-- Fundamental theorem of calculus: if `f : ℝ → E` is integrable on `a..b` and `f x` has a finite -limit `c` almost surely as `x` tends to `a` from the right or from the left, then the right -(resp., left) derivative of `u ↦ ∫ x in u..b, f x` at `a` equals `-c`. -/ -lemma deriv_within_integral_of_tendsto_ae_left - (hf : interval_integrable f volume a b) {s t : set ℝ} [FTC_filter a (𝓝[s] a) (𝓝[t] a)] - (hmeas : strongly_measurable_at_filter f (𝓝[t] a)) (ha : tendsto f (𝓝[t] a ⊓ volume.ae) (𝓝 c)) - (hs : unique_diff_within_at ℝ s a . unique_diff_within_at_Ici_Iic_univ) : - deriv_within (λ u, ∫ x in u..b, f x) s a = -c := -(integral_has_deriv_within_at_of_tendsto_ae_left hf hmeas ha).deriv_within hs - -/-- Fundamental theorem of calculus: if `f : ℝ → E` is integrable on `a..b` and `f x` is continuous -on the right or on the left at `a`, then the right (resp., left) derivative of -`u ↦ ∫ x in u..b, f x` at `a` equals `-f a`. -/ -lemma deriv_within_integral_left - (hf : interval_integrable f volume a b) {s t : set ℝ} [FTC_filter a (𝓝[s] a) (𝓝[t] a)] - (hmeas : strongly_measurable_at_filter f (𝓝[t] a)) (ha : continuous_within_at f t a) - (hs : unique_diff_within_at ℝ s a . unique_diff_within_at_Ici_Iic_univ) : - deriv_within (λ u, ∫ x in u..b, f x) s a = -f a := -(integral_has_deriv_within_at_left hf hmeas ha).deriv_within hs - -/-- The integral of a continuous function is differentiable on a real set `s`. -/ -theorem differentiable_on_integral_of_continuous {s : set ℝ} - (hintg : ∀ x ∈ s, interval_integrable f volume a x) (hcont : continuous f) : - differentiable_on ℝ (λ u, ∫ x in a..u, f x) s := -λ y hy, (integral_has_deriv_at_right (hintg y hy) - hcont.ae_strongly_measurable.strongly_measurable_at_filter - hcont.continuous_at) .differentiable_at.differentiable_within_at - -/-! -### Fundamental theorem of calculus, part 2 - -This section contains theorems pertaining to FTC-2 for interval integrals, i.e., the assertion -that `∫ x in a..b, f' x = f b - f a` under suitable assumptions. - -The most classical version of this theorem assumes that `f'` is continuous. However, this is -unnecessarily strong: the result holds if `f'` is just integrable. We prove the strong version, -following [Rudin, *Real and Complex Analysis* (Theorem 7.21)][rudin2006real]. The proof is first -given for real-valued functions, and then deduced for functions with a general target space. For -a real-valued function `g`, it suffices to show that `g b - g a ≤ (∫ x in a..b, g' x) + ε` for all -positive `ε`. To prove this, choose a lower-semicontinuous function `G'` with `g' < G'` and with -integral close to that of `g'` (its existence is guaranteed by the Vitali-Carathéodory theorem). -It satisfies `g t - g a ≤ ∫ x in a..t, G' x` for all `t ∈ [a, b]`: this inequality holds at `a`, -and if it holds at `t` then it holds for `u` close to `t` on its right, as the left hand side -increases by `g u - g t ∼ (u -t) g' t`, while the right hand side increases by -`∫ x in t..u, G' x` which is roughly at least `∫ x in t..u, G' t = (u - t) G' t`, by lower -semicontinuity. As `g' t < G' t`, this gives the conclusion. One can therefore push progressively -this inequality to the right until the point `b`, where it gives the desired conclusion. --/ - -variables {g' g φ : ℝ → ℝ} - -/-- Hard part of FTC-2 for integrable derivatives, real-valued functions: one has -`g b - g a ≤ ∫ y in a..b, g' y` when `g'` is integrable. -Auxiliary lemma in the proof of `integral_eq_sub_of_has_deriv_right_of_le`. -We give the slightly more general version that `g b - g a ≤ ∫ y in a..b, φ y` when `g' ≤ φ` and -`φ` is integrable (even if `g'` is not known to be integrable). -Version assuming that `g` is differentiable on `[a, b)`. -/ -lemma sub_le_integral_of_has_deriv_right_of_le_Ico (hab : a ≤ b) (hcont : continuous_on g (Icc a b)) - (hderiv : ∀ x ∈ Ico a b, has_deriv_within_at g (g' x) (Ioi x) x) - (φint : integrable_on φ (Icc a b)) (hφg : ∀ x ∈ Ico a b, g' x ≤ φ x) : - g b - g a ≤ ∫ y in a..b, φ y := -begin - refine le_of_forall_pos_le_add (λ ε εpos, _), - -- Bound from above `g'` by a lower-semicontinuous function `G'`. - rcases exists_lt_lower_semicontinuous_integral_lt φ φint εpos with - ⟨G', f_lt_G', G'cont, G'int, G'lt_top, hG'⟩, - -- we will show by "induction" that `g t - g a ≤ ∫ u in a..t, G' u` for all `t ∈ [a, b]`. - set s := {t | g t - g a ≤ ∫ u in a..t, (G' u).to_real} ∩ Icc a b, - -- the set `s` of points where this property holds is closed. - have s_closed : is_closed s, - { have : continuous_on (λ t, (g t - g a, ∫ u in a..t, (G' u).to_real)) (Icc a b), - { rw ← uIcc_of_le hab at G'int ⊢ hcont, - exact (hcont.sub continuous_on_const).prod (continuous_on_primitive_interval G'int) }, - simp only [s, inter_comm], - exact this.preimage_closed_of_closed is_closed_Icc order_closed_topology.is_closed_le' }, - have main : Icc a b ⊆ {t | g t - g a ≤ ∫ u in a..t, (G' u).to_real }, - { -- to show that the set `s` is all `[a, b]`, it suffices to show that any point `t` in `s` - -- with `t < b` admits another point in `s` slightly to its right - -- (this is a sort of real induction). - apply s_closed.Icc_subset_of_forall_exists_gt - (by simp only [integral_same, mem_set_of_eq, sub_self]) (λ t ht v t_lt_v, _), - obtain ⟨y, g'_lt_y', y_lt_G'⟩ : ∃ (y : ℝ), (g' t : ereal) < y ∧ (y : ereal) < G' t := - ereal.lt_iff_exists_real_btwn.1 ((ereal.coe_le_coe_iff.2 (hφg t ht.2)).trans_lt (f_lt_G' t)), - -- bound from below the increase of `∫ x in a..u, G' x` on the right of `t`, using the lower - -- semicontinuity of `G'`. - have I1 : ∀ᶠ u in 𝓝[>] t, (u - t) * y ≤ ∫ w in t..u, (G' w).to_real, - { have B : ∀ᶠ u in 𝓝 t, (y : ereal) < G' u := - G'cont.lower_semicontinuous_at _ _ y_lt_G', - rcases mem_nhds_iff_exists_Ioo_subset.1 B with ⟨m, M, ⟨hm, hM⟩, H⟩, - have : Ioo t (min M b) ∈ 𝓝[>] t := mem_nhds_within_Ioi_iff_exists_Ioo_subset.2 - ⟨min M b, by simp only [hM, ht.right.right, lt_min_iff, mem_Ioi, and_self], subset.refl _⟩, - filter_upwards [this] with u hu, - have I : Icc t u ⊆ Icc a b := Icc_subset_Icc ht.2.1 (hu.2.le.trans (min_le_right _ _)), - calc (u - t) * y = ∫ v in Icc t u, y : - by simp only [hu.left.le, measure_theory.integral_const, algebra.id.smul_eq_mul, sub_nonneg, - measurable_set.univ, real.volume_Icc, measure.restrict_apply, univ_inter, - ennreal.to_real_of_real] - ... ≤ ∫ w in t..u, (G' w).to_real : - begin - rw [interval_integral.integral_of_le hu.1.le, ← integral_Icc_eq_integral_Ioc], - apply set_integral_mono_ae_restrict, - { simp only [integrable_on_const, real.volume_Icc, ennreal.of_real_lt_top, or_true] }, - { exact integrable_on.mono_set G'int I }, - { have C1 : ∀ᵐ (x : ℝ) ∂volume.restrict (Icc t u), G' x < ∞ := - ae_mono (measure.restrict_mono I le_rfl) G'lt_top, - have C2 : ∀ᵐ (x : ℝ) ∂volume.restrict (Icc t u), x ∈ Icc t u := - ae_restrict_mem measurable_set_Icc, - filter_upwards [C1, C2] with x G'x hx, - apply ereal.coe_le_coe_iff.1, - have : x ∈ Ioo m M, by simp only [hm.trans_le hx.left, - (hx.right.trans_lt hu.right).trans_le (min_le_left M b), mem_Ioo, and_self], - convert le_of_lt (H this), - exact ereal.coe_to_real G'x.ne (ne_bot_of_gt (f_lt_G' x)) } - end }, - -- bound from above the increase of `g u - g a` on the right of `t`, using the derivative at `t` - have I2 : ∀ᶠ u in 𝓝[>] t, g u - g t ≤ (u - t) * y, - { have g'_lt_y : g' t < y := ereal.coe_lt_coe_iff.1 g'_lt_y', - filter_upwards [(hderiv t ⟨ht.2.1, ht.2.2⟩).limsup_slope_le' - (not_mem_Ioi.2 le_rfl) g'_lt_y, self_mem_nhds_within] with u hu t_lt_u, - have := mul_le_mul_of_nonneg_left hu.le (sub_pos.2 t_lt_u).le, - rwa [← smul_eq_mul, sub_smul_slope] at this }, - -- combine the previous two bounds to show that `g u - g a` increases less quickly than - -- `∫ x in a..u, G' x`. - have I3 : ∀ᶠ u in 𝓝[>] t, g u - g t ≤ ∫ w in t..u, (G' w).to_real, - { filter_upwards [I1, I2] with u hu1 hu2 using hu2.trans hu1, }, - have I4 : ∀ᶠ u in 𝓝[>] t, u ∈ Ioc t (min v b), - { refine mem_nhds_within_Ioi_iff_exists_Ioc_subset.2 ⟨min v b, _, subset.refl _⟩, - simp only [lt_min_iff, mem_Ioi], - exact ⟨t_lt_v, ht.2.2⟩ }, - -- choose a point `x` slightly to the right of `t` which satisfies the above bound - rcases (I3.and I4).exists with ⟨x, hx, h'x⟩, - -- we check that it belongs to `s`, essentially by construction - refine ⟨x, _, Ioc_subset_Ioc le_rfl (min_le_left _ _) h'x⟩, - calc g x - g a = (g t - g a) + (g x - g t) : by abel - ... ≤ (∫ w in a..t, (G' w).to_real) + ∫ w in t..x, (G' w).to_real : add_le_add ht.1 hx - ... = ∫ w in a..x, (G' w).to_real : - begin - apply integral_add_adjacent_intervals, - { rw interval_integrable_iff_integrable_Ioc_of_le ht.2.1, - exact integrable_on.mono_set G'int - (Ioc_subset_Icc_self.trans (Icc_subset_Icc le_rfl ht.2.2.le)) }, - { rw interval_integrable_iff_integrable_Ioc_of_le h'x.1.le, - apply integrable_on.mono_set G'int, - refine Ioc_subset_Icc_self.trans (Icc_subset_Icc ht.2.1 (h'x.2.trans (min_le_right _ _))) } - end }, - -- now that we know that `s` contains `[a, b]`, we get the desired result by applying this to `b`. - calc g b - g a ≤ ∫ y in a..b, (G' y).to_real : main (right_mem_Icc.2 hab) - ... ≤ (∫ y in a..b, φ y) + ε : - begin - convert hG'.le; - { rw interval_integral.integral_of_le hab, - simp only [integral_Icc_eq_integral_Ioc', real.volume_singleton] }, - end -end - -/-- Hard part of FTC-2 for integrable derivatives, real-valued functions: one has -`g b - g a ≤ ∫ y in a..b, g' y` when `g'` is integrable. -Auxiliary lemma in the proof of `integral_eq_sub_of_has_deriv_right_of_le`. -We give the slightly more general version that `g b - g a ≤ ∫ y in a..b, φ y` when `g' ≤ φ` and -`φ` is integrable (even if `g'` is not known to be integrable). -Version assuming that `g` is differentiable on `(a, b)`. -/ -lemma sub_le_integral_of_has_deriv_right_of_le (hab : a ≤ b) - (hcont : continuous_on g (Icc a b)) - (hderiv : ∀ x ∈ Ioo a b, has_deriv_within_at g (g' x) (Ioi x) x) - (φint : integrable_on φ (Icc a b)) (hφg : ∀ x ∈ Ioo a b, g' x ≤ φ x) : - g b - g a ≤ ∫ y in a..b, φ y := -begin - -- This follows from the version on a closed-open interval (applied to `[t, b)` for `t` close to - -- `a`) and a continuity argument. - obtain rfl|a_lt_b := hab.eq_or_lt, { simp }, - set s := {t | g b - g t ≤ ∫ u in t..b, φ u} ∩ Icc a b, - have s_closed : is_closed s, - { have : continuous_on (λ t, (g b - g t, ∫ u in t..b, φ u)) (Icc a b), - { rw ← uIcc_of_le hab at ⊢ hcont φint, - exact (continuous_on_const.sub hcont).prod (continuous_on_primitive_interval_left φint) }, - simp only [s, inter_comm], - exact this.preimage_closed_of_closed is_closed_Icc is_closed_le_prod, }, - have A : closure (Ioc a b) ⊆ s, - { apply s_closed.closure_subset_iff.2, - assume t ht, - refine ⟨_, ⟨ht.1.le, ht.2⟩⟩, - exact sub_le_integral_of_has_deriv_right_of_le_Ico ht.2 - (hcont.mono (Icc_subset_Icc ht.1.le le_rfl)) - (λ x hx, hderiv x ⟨ht.1.trans_le hx.1, hx.2⟩) - (φint.mono_set (Icc_subset_Icc ht.1.le le_rfl)) - (λ x hx, hφg x ⟨ht.1.trans_le hx.1, hx.2⟩) }, - rw closure_Ioc a_lt_b.ne at A, - exact (A (left_mem_Icc.2 hab)).1, -end - -/-- Auxiliary lemma in the proof of `integral_eq_sub_of_has_deriv_right_of_le`. -/ -lemma integral_le_sub_of_has_deriv_right_of_le (hab : a ≤ b) - (hcont : continuous_on g (Icc a b)) - (hderiv : ∀ x ∈ Ioo a b, has_deriv_within_at g (g' x) (Ioi x) x) - (φint : integrable_on φ (Icc a b)) (hφg : ∀ x ∈ Ioo a b, φ x ≤ g' x) : - ∫ y in a..b, φ y ≤ g b - g a := -begin - rw ← neg_le_neg_iff, - convert sub_le_integral_of_has_deriv_right_of_le hab hcont.neg (λ x hx, (hderiv x hx).neg) - φint.neg (λ x hx, neg_le_neg (hφg x hx)), - { abel }, - { simp only [← integral_neg], refl }, -end - -/-- Auxiliary lemma in the proof of `integral_eq_sub_of_has_deriv_right_of_le`: real version -/ -lemma integral_eq_sub_of_has_deriv_right_of_le_real (hab : a ≤ b) - (hcont : continuous_on g (Icc a b)) - (hderiv : ∀ x ∈ Ioo a b, has_deriv_within_at g (g' x) (Ioi x) x) - (g'int : integrable_on g' (Icc a b)) : - ∫ y in a..b, g' y = g b - g a := -le_antisymm - (integral_le_sub_of_has_deriv_right_of_le hab hcont hderiv g'int (λ x hx, le_rfl)) - (sub_le_integral_of_has_deriv_right_of_le hab hcont hderiv g'int (λ x hx, le_rfl)) - -variable {f' : ℝ → E} - -/-- **Fundamental theorem of calculus-2**: If `f : ℝ → E` is continuous on `[a, b]` (where `a ≤ b`) - and has a right derivative at `f' x` for all `x` in `(a, b)`, and `f'` is integrable on `[a, b]`, - then `∫ y in a..b, f' y` equals `f b - f a`. -/ -theorem integral_eq_sub_of_has_deriv_right_of_le (hab : a ≤ b) (hcont : continuous_on f (Icc a b)) - (hderiv : ∀ x ∈ Ioo a b, has_deriv_within_at f (f' x) (Ioi x) x) - (f'int : interval_integrable f' volume a b) : - ∫ y in a..b, f' y = f b - f a := -begin - refine (normed_space.eq_iff_forall_dual_eq ℝ).2 (λ g, _), - rw [← g.interval_integral_comp_comm f'int, g.map_sub], - exact integral_eq_sub_of_has_deriv_right_of_le_real hab (g.continuous.comp_continuous_on hcont) - (λ x hx, g.has_fderiv_at.comp_has_deriv_within_at x (hderiv x hx)) - (g.integrable_comp ((interval_integrable_iff_integrable_Icc_of_le hab).1 f'int)) -end - -/-- Fundamental theorem of calculus-2: If `f : ℝ → E` is continuous on `[a, b]` and - has a right derivative at `f' x` for all `x` in `[a, b)`, and `f'` is integrable on `[a, b]` then - `∫ y in a..b, f' y` equals `f b - f a`. -/ -theorem integral_eq_sub_of_has_deriv_right (hcont : continuous_on f (uIcc a b)) - (hderiv : ∀ x ∈ Ioo (min a b) (max a b), has_deriv_within_at f (f' x) (Ioi x) x) - (hint : interval_integrable f' volume a b) : - ∫ y in a..b, f' y = f b - f a := -begin - cases le_total a b with hab hab, - { simp only [uIcc_of_le, min_eq_left, max_eq_right, hab] at hcont hderiv hint, - apply integral_eq_sub_of_has_deriv_right_of_le hab hcont hderiv hint }, - { simp only [uIcc_of_ge, min_eq_right, max_eq_left, hab] at hcont hderiv, - rw [integral_symm, integral_eq_sub_of_has_deriv_right_of_le hab hcont hderiv hint.symm, - neg_sub] } -end - -/-- Fundamental theorem of calculus-2: If `f : ℝ → E` is continuous on `[a, b]` (where `a ≤ b`) and - has a derivative at `f' x` for all `x` in `(a, b)`, and `f'` is integrable on `[a, b]`, then - `∫ y in a..b, f' y` equals `f b - f a`. -/ -theorem integral_eq_sub_of_has_deriv_at_of_le (hab : a ≤ b) - (hcont : continuous_on f (Icc a b)) - (hderiv : ∀ x ∈ Ioo a b, has_deriv_at f (f' x) x) (hint : interval_integrable f' volume a b) : - ∫ y in a..b, f' y = f b - f a := -integral_eq_sub_of_has_deriv_right_of_le hab hcont (λ x hx, (hderiv x hx).has_deriv_within_at) hint - -/-- Fundamental theorem of calculus-2: If `f : ℝ → E` has a derivative at `f' x` for all `x` in - `[a, b]` and `f'` is integrable on `[a, b]`, then `∫ y in a..b, f' y` equals `f b - f a`. -/ -theorem integral_eq_sub_of_has_deriv_at - (hderiv : ∀ x ∈ uIcc a b, has_deriv_at f (f' x) x) - (hint : interval_integrable f' volume a b) : - ∫ y in a..b, f' y = f b - f a := -integral_eq_sub_of_has_deriv_right (has_deriv_at.continuous_on hderiv) - (λ x hx, (hderiv _ (mem_Icc_of_Ioo hx)).has_deriv_within_at) hint - -theorem integral_eq_sub_of_has_deriv_at_of_tendsto (hab : a < b) {fa fb} - (hderiv : ∀ x ∈ Ioo a b, has_deriv_at f (f' x) x) (hint : interval_integrable f' volume a b) - (ha : tendsto f (𝓝[>] a) (𝓝 fa)) (hb : tendsto f (𝓝[<] b) (𝓝 fb)) : - ∫ y in a..b, f' y = fb - fa := -begin - set F : ℝ → E := update (update f a fa) b fb, - have Fderiv : ∀ x ∈ Ioo a b, has_deriv_at F (f' x) x, - { refine λ x hx, (hderiv x hx).congr_of_eventually_eq _, - filter_upwards [Ioo_mem_nhds hx.1 hx.2] with _ hy, simp only [F], - rw [update_noteq hy.2.ne, update_noteq hy.1.ne'], }, - have hcont : continuous_on F (Icc a b), - { rw [continuous_on_update_iff, continuous_on_update_iff, Icc_diff_right, Ico_diff_left], - refine ⟨⟨λ z hz, (hderiv z hz).continuous_at.continuous_within_at, _⟩, _⟩, - { exact λ _, ha.mono_left (nhds_within_mono _ Ioo_subset_Ioi_self) }, - { rintro -, - refine (hb.congr' _).mono_left (nhds_within_mono _ Ico_subset_Iio_self), - filter_upwards [Ioo_mem_nhds_within_Iio (right_mem_Ioc.2 hab)] - with _ hz using (update_noteq hz.1.ne' _ _).symm } }, - simpa [F, hab.ne, hab.ne'] using integral_eq_sub_of_has_deriv_at_of_le hab.le hcont Fderiv hint -end - -/-- Fundamental theorem of calculus-2: If `f : ℝ → E` is differentiable at every `x` in `[a, b]` and - its derivative is integrable on `[a, b]`, then `∫ y in a..b, deriv f y` equals `f b - f a`. -/ -theorem integral_deriv_eq_sub (hderiv : ∀ x ∈ uIcc a b, differentiable_at ℝ f x) - (hint : interval_integrable (deriv f) volume a b) : - ∫ y in a..b, deriv f y = f b - f a := -integral_eq_sub_of_has_deriv_at (λ x hx, (hderiv x hx).has_deriv_at) hint - -theorem integral_deriv_eq_sub' (f) (hderiv : deriv f = f') - (hdiff : ∀ x ∈ uIcc a b, differentiable_at ℝ f x) - (hcont : continuous_on f' (uIcc a b)) : - ∫ y in a..b, f' y = f b - f a := -begin - rw [← hderiv, integral_deriv_eq_sub hdiff], - rw hderiv, - exact hcont.interval_integrable -end - -/-! -### Automatic integrability for nonnegative derivatives --/ - -/-- When the right derivative of a function is nonnegative, then it is automatically integrable. -/ -lemma integrable_on_deriv_right_of_nonneg (hcont : continuous_on g (Icc a b)) - (hderiv : ∀ x ∈ Ioo a b, has_deriv_within_at g (g' x) (Ioi x) x) - (g'pos : ∀ x ∈ Ioo a b, 0 ≤ g' x) : - integrable_on g' (Ioc a b) := -begin - by_cases hab : a < b, swap, - { simp [Ioc_eq_empty hab] }, - rw integrable_on_Ioc_iff_integrable_on_Ioo, - have meas_g' : ae_measurable g' (volume.restrict (Ioo a b)), - { apply (ae_measurable_deriv_within_Ioi g _).congr, - refine (ae_restrict_mem measurable_set_Ioo).mono (λ x hx, _), - exact (hderiv x hx).deriv_within (unique_diff_within_at_Ioi _) }, - suffices H : ∫⁻ x in Ioo a b, ‖g' x‖₊ ≤ ennreal.of_real (g b - g a), - from ⟨meas_g'.ae_strongly_measurable, H.trans_lt ennreal.of_real_lt_top⟩, - by_contra' H, - obtain ⟨f, fle, fint, hf⟩ : - ∃ (f : simple_func ℝ ℝ≥0), (∀ x, f x ≤ ‖g' x‖₊) ∧ ∫⁻ (x : ℝ) in Ioo a b, f x < ∞ - ∧ ennreal.of_real (g b - g a) < ∫⁻ (x : ℝ) in Ioo a b, f x := - exists_lt_lintegral_simple_func_of_lt_lintegral H, - let F : ℝ → ℝ := coe ∘ f, - have intF : integrable_on F (Ioo a b), - { refine ⟨f.measurable.coe_nnreal_real.ae_strongly_measurable, _⟩, - simpa only [has_finite_integral, nnreal.nnnorm_eq] using fint }, - have A : ∫⁻ (x : ℝ) in Ioo a b, f x = ennreal.of_real (∫ x in Ioo a b, F x) := - lintegral_coe_eq_integral _ intF, - rw A at hf, - have B : ∫ (x : ℝ) in Ioo a b, F x ≤ g b - g a, - { rw [← integral_Ioc_eq_integral_Ioo, ← interval_integral.integral_of_le hab.le], - apply integral_le_sub_of_has_deriv_right_of_le hab.le hcont hderiv _ (λ x hx, _), - { rwa integrable_on_Icc_iff_integrable_on_Ioo }, - { convert nnreal.coe_le_coe.2 (fle x), - simp only [real.norm_of_nonneg (g'pos x hx), coe_nnnorm] } }, - exact lt_irrefl _ (hf.trans_le (ennreal.of_real_le_of_real B)), -end - -/-- When the derivative of a function is nonnegative, then it is automatically integrable, -Ioc version. -/ -lemma integrable_on_deriv_of_nonneg (hcont : continuous_on g (Icc a b)) - (hderiv : ∀ x ∈ Ioo a b, has_deriv_at g (g' x) x) - (g'pos : ∀ x ∈ Ioo a b, 0 ≤ g' x) : - integrable_on g' (Ioc a b) := -integrable_on_deriv_right_of_nonneg hcont (λ x hx, (hderiv x hx).has_deriv_within_at) g'pos - -/-- When the derivative of a function is nonnegative, then it is automatically integrable, -interval version. -/ -theorem interval_integrable_deriv_of_nonneg (hcont : continuous_on g (uIcc a b)) - (hderiv : ∀ x ∈ Ioo (min a b) (max a b), has_deriv_at g (g' x) x) - (hpos : ∀ x ∈ Ioo (min a b) (max a b), 0 ≤ g' x) : - interval_integrable g' volume a b := -begin - cases le_total a b with hab hab, - { simp only [uIcc_of_le, min_eq_left, max_eq_right, hab, interval_integrable, - hab, Ioc_eq_empty_of_le, integrable_on_empty, and_true] at hcont hderiv hpos ⊢, - exact integrable_on_deriv_of_nonneg hcont hderiv hpos, }, - { simp only [uIcc_of_ge, min_eq_right, max_eq_left, hab, interval_integrable, - Ioc_eq_empty_of_le, integrable_on_empty, true_and] at hcont hderiv hpos ⊢, - exact integrable_on_deriv_of_nonneg hcont hderiv hpos } -end - -/-! -### Integration by parts --/ - -section parts - -variables [normed_ring A] [normed_algebra ℝ A] [complete_space A] - -theorem integral_deriv_mul_eq_sub {u v u' v' : ℝ → A} - (hu : ∀ x ∈ uIcc a b, has_deriv_at u (u' x) x) - (hv : ∀ x ∈ uIcc a b, has_deriv_at v (v' x) x) - (hu' : interval_integrable u' volume a b) (hv' : interval_integrable v' volume a b) : - ∫ x in a..b, u' x * v x + u x * v' x = u b * v b - u a * v a := -integral_eq_sub_of_has_deriv_at (λ x hx, (hu x hx).mul (hv x hx)) $ - (hu'.mul_continuous_on (has_deriv_at.continuous_on hv)).add - (hv'.continuous_on_mul ((has_deriv_at.continuous_on hu))) - -theorem integral_mul_deriv_eq_deriv_mul {u v u' v' : ℝ → A} - (hu : ∀ x ∈ uIcc a b, has_deriv_at u (u' x) x) - (hv : ∀ x ∈ uIcc a b, has_deriv_at v (v' x) x) - (hu' : interval_integrable u' volume a b) (hv' : interval_integrable v' volume a b) : - ∫ x in a..b, u x * v' x = u b * v b - u a * v a - ∫ x in a..b, u' x * v x := -begin - rw [← integral_deriv_mul_eq_sub hu hv hu' hv', ← integral_sub], - { exact integral_congr (λ x hx, by simp only [add_sub_cancel']) }, - { exact ((hu'.mul_continuous_on (has_deriv_at.continuous_on hv)).add - (hv'.continuous_on_mul (has_deriv_at.continuous_on hu))) }, - { exact hu'.mul_continuous_on (has_deriv_at.continuous_on hv) }, -end - -end parts - -/-! -### Integration by substitution / Change of variables --/ - -section smul - -/-- -Change of variables, general form. If `f` is continuous on `[a, b]` and has -right-derivative `f'` in `(a, b)`, `g` is continuous on `f '' (a, b)` and integrable on -`f '' [a, b]`, and `f' x • (g ∘ f) x` is integrable on `[a, b]`, -then we can substitute `u = f x` to get `∫ x in a..b, f' x • (g ∘ f) x = ∫ u in f a..f b, g u`. --/ -theorem integral_comp_smul_deriv''' {f f' : ℝ → ℝ} {g : ℝ → E} - (hf : continuous_on f [a, b]) - (hff' : ∀ x ∈ Ioo (min a b) (max a b), has_deriv_within_at f (f' x) (Ioi x) x) - (hg_cont : continuous_on g (f '' Ioo (min a b) (max a b))) - (hg1 : integrable_on g (f '' [a, b]) ) - (hg2 : integrable_on (λ x, f'(x) • (g ∘ f) x) [a, b]) : - ∫ x in a..b, f' x • (g ∘ f) x = ∫ u in f a..f b, g u := -begin - rw [hf.image_uIcc, ←interval_integrable_iff'] at hg1, - have h_cont : continuous_on (λ u, ∫ t in f a..f u, g t) [a, b], - { refine (continuous_on_primitive_interval' hg1 _).comp hf _, - { rw ← hf.image_uIcc, exact mem_image_of_mem f left_mem_uIcc }, - { rw ← hf.image_uIcc, exact maps_to_image _ _ } }, - have h_der : ∀ x ∈ Ioo (min a b) (max a b), has_deriv_within_at - (λ u, ∫ t in f a..f u, g t) (f' x • ((g ∘ f) x)) (Ioi x) x, - { intros x hx, - obtain ⟨c, hc⟩ := nonempty_Ioo.mpr hx.1, - obtain ⟨d, hd⟩ := nonempty_Ioo.mpr hx.2, - have cdsub : [c, d] ⊆ Ioo (min a b) (max a b), - { rw uIcc_of_le (hc.2.trans hd.1).le, exact Icc_subset_Ioo hc.1 hd.2 }, - replace hg_cont := hg_cont.mono (image_subset f cdsub), - let J := [Inf (f '' [c, d]), Sup (f '' [c, d])], - have hJ : f '' [c, d] = J := (hf.mono (cdsub.trans Ioo_subset_Icc_self)).image_uIcc, - rw hJ at hg_cont, - have h2x : f x ∈ J, { rw ←hJ, exact mem_image_of_mem _ (mem_uIcc_of_le hc.2.le hd.1.le), }, - have h2g : interval_integrable g volume (f a) (f x), - { refine hg1.mono_set _, - rw ←hf.image_uIcc, - exact hf.surj_on_uIcc left_mem_uIcc (Ioo_subset_Icc_self hx) }, - have h3g := hg_cont.strongly_measurable_at_filter_nhds_within measurable_set_Icc (f x), - haveI : fact (f x ∈ J) := ⟨h2x⟩, - have : has_deriv_within_at (λ u, ∫ x in f a..u, g x) (g (f x)) J (f x) := - interval_integral.integral_has_deriv_within_at_right h2g h3g (hg_cont (f x) h2x), - refine (this.scomp x ((hff' x hx).Ioo_of_Ioi hd.1) _).Ioi_of_Ioo hd.1, - rw ←hJ, - refine (maps_to_image _ _).mono _ subset.rfl, - exact Ioo_subset_Icc_self.trans ((Icc_subset_Icc_left hc.2.le).trans Icc_subset_uIcc) }, - rw ←interval_integrable_iff' at hg2, - simp_rw [integral_eq_sub_of_has_deriv_right h_cont h_der hg2, integral_same, sub_zero], -end - -/-- -Change of variables for continuous integrands. If `f` is continuous on `[a, b]` and has -continuous right-derivative `f'` in `(a, b)`, and `g` is continuous on `f '' [a, b]` then we can -substitute `u = f x` to get `∫ x in a..b, f' x • (g ∘ f) x = ∫ u in f a..f b, g u`. --/ -theorem integral_comp_smul_deriv'' {f f' : ℝ → ℝ} {g : ℝ → E} - (hf : continuous_on f [a, b]) - (hff' : ∀ x ∈ Ioo (min a b) (max a b), has_deriv_within_at f (f' x) (Ioi x) x) - (hf' : continuous_on f' [a, b]) - (hg : continuous_on g (f '' [a, b])) : - ∫ x in a..b, f' x • (g ∘ f) x = ∫ u in f a..f b, g u := -begin - refine integral_comp_smul_deriv''' hf hff' - (hg.mono $ image_subset _ Ioo_subset_Icc_self) _ - (hf'.smul (hg.comp hf $ subset_preimage_image f _)).integrable_on_Icc, - rw hf.image_uIcc at hg ⊢, - exact hg.integrable_on_Icc, -end - -/-- -Change of variables. If `f` is has continuous derivative `f'` on `[a, b]`, -and `g` is continuous on `f '' [a, b]`, then we can substitute `u = f x` to get -`∫ x in a..b, f' x • (g ∘ f) x = ∫ u in f a..f b, g u`. -Compared to `interval_integral.integral_comp_smul_deriv` we only require that `g` is continuous on -`f '' [a, b]`. --/ -theorem integral_comp_smul_deriv' {f f' : ℝ → ℝ} {g : ℝ → E} - (h : ∀ x ∈ uIcc a b, has_deriv_at f (f' x) x) - (h' : continuous_on f' (uIcc a b)) (hg : continuous_on g (f '' [a, b])) : - ∫ x in a..b, f' x • (g ∘ f) x = ∫ x in f a..f b, g x := -integral_comp_smul_deriv'' (λ x hx, (h x hx).continuous_at.continuous_within_at) - (λ x hx, (h x $ Ioo_subset_Icc_self hx).has_deriv_within_at) h' hg - -/-- -Change of variables, most common version. If `f` is has continuous derivative `f'` on `[a, b]`, -and `g` is continuous, then we can substitute `u = f x` to get -`∫ x in a..b, f' x • (g ∘ f) x = ∫ u in f a..f b, g u`. --/ -theorem integral_comp_smul_deriv {f f' : ℝ → ℝ} {g : ℝ → E} - (h : ∀ x ∈ uIcc a b, has_deriv_at f (f' x) x) - (h' : continuous_on f' (uIcc a b)) (hg : continuous g) : - ∫ x in a..b, f' x • (g ∘ f) x = ∫ x in f a..f b, g x := -integral_comp_smul_deriv' h h' hg.continuous_on - -theorem integral_deriv_comp_smul_deriv' {f f' : ℝ → ℝ} {g g' : ℝ → E} - (hf : continuous_on f [a, b]) - (hff' : ∀ x ∈ Ioo (min a b) (max a b), has_deriv_within_at f (f' x) (Ioi x) x) - (hf' : continuous_on f' [a, b]) - (hg : continuous_on g [f a, f b]) - (hgg' : ∀ x ∈ Ioo (min (f a) (f b)) (max (f a) (f b)), has_deriv_within_at g (g' x) (Ioi x) x) - (hg' : continuous_on g' (f '' [a, b])) : - ∫ x in a..b, f' x • (g' ∘ f) x = (g ∘ f) b - (g ∘ f) a := -begin - rw [integral_comp_smul_deriv'' hf hff' hf' hg', - integral_eq_sub_of_has_deriv_right hg hgg' (hg'.mono _).interval_integrable], - exact intermediate_value_uIcc hf -end - -theorem integral_deriv_comp_smul_deriv {f f' : ℝ → ℝ} {g g' : ℝ → E} - (hf : ∀ x ∈ uIcc a b, has_deriv_at f (f' x) x) - (hg : ∀ x ∈ uIcc a b, has_deriv_at g (g' (f x)) (f x)) - (hf' : continuous_on f' (uIcc a b)) (hg' : continuous g') : - ∫ x in a..b, f' x • (g' ∘ f) x = (g ∘ f) b - (g ∘ f) a := -integral_eq_sub_of_has_deriv_at (λ x hx, (hg x hx).scomp x $ hf x hx) - (hf'.smul (hg'.comp_continuous_on $ has_deriv_at.continuous_on hf)).interval_integrable - -end smul -section mul - -/-- -Change of variables, general form for scalar functions. If `f` is continuous on `[a, b]` and has -continuous right-derivative `f'` in `(a, b)`, `g` is continuous on `f '' (a, b)` and integrable on -`f '' [a, b]`, and `(g ∘ f) x * f' x` is integrable on `[a, b]`, then we can substitute `u = f x` -to get `∫ x in a..b, (g ∘ f) x * f' x = ∫ u in f a..f b, g u`. --/ -theorem integral_comp_mul_deriv''' {a b : ℝ} {f f' : ℝ → ℝ} {g : ℝ → ℝ} - (hf : continuous_on f [a, b]) - (hff' : ∀ x ∈ Ioo (min a b) (max a b), has_deriv_within_at f (f' x) (Ioi x) x) - (hg_cont : continuous_on g (f '' Ioo (min a b) (max a b))) - (hg1 : integrable_on g (f '' [a, b]) ) - (hg2 : integrable_on (λ x, (g ∘ f) x * f' x) [a, b]) : - ∫ x in a..b, (g ∘ f) x * f' x = ∫ u in f a..f b, g u := -begin - have hg2' : integrable_on (λ x, f' x • (g ∘ f) x) [a, b] := by simpa [mul_comm] using hg2, - simpa [mul_comm] using integral_comp_smul_deriv''' hf hff' hg_cont hg1 hg2', -end - -/-- -Change of variables for continuous integrands. If `f` is continuous on `[a, b]` and has -continuous right-derivative `f'` in `(a, b)`, and `g` is continuous on `f '' [a, b]` then we can -substitute `u = f x` to get `∫ x in a..b, (g ∘ f) x * f' x = ∫ u in f a..f b, g u`. --/ -theorem integral_comp_mul_deriv'' {f f' g : ℝ → ℝ} - (hf : continuous_on f [a, b]) - (hff' : ∀ x ∈ Ioo (min a b) (max a b), has_deriv_within_at f (f' x) (Ioi x) x) - (hf' : continuous_on f' [a, b]) - (hg : continuous_on g (f '' [a, b])) : - ∫ x in a..b, (g ∘ f) x * f' x = ∫ u in f a..f b, g u := -by simpa [mul_comm] using integral_comp_smul_deriv'' hf hff' hf' hg - -/-- -Change of variables. If `f` is has continuous derivative `f'` on `[a, b]`, -and `g` is continuous on `f '' [a, b]`, then we can substitute `u = f x` to get -`∫ x in a..b, (g ∘ f) x * f' x = ∫ u in f a..f b, g u`. -Compared to `interval_integral.integral_comp_mul_deriv` we only require that `g` is continuous on -`f '' [a, b]`. --/ -theorem integral_comp_mul_deriv' {f f' g : ℝ → ℝ} - (h : ∀ x ∈ uIcc a b, has_deriv_at f (f' x) x) - (h' : continuous_on f' (uIcc a b)) (hg : continuous_on g (f '' [a, b])) : - ∫ x in a..b, (g ∘ f) x * f' x = ∫ x in f a..f b, g x := -by simpa [mul_comm] using integral_comp_smul_deriv' h h' hg - -/-- -Change of variables, most common version. If `f` is has continuous derivative `f'` on `[a, b]`, -and `g` is continuous, then we can substitute `u = f x` to get -`∫ x in a..b, (g ∘ f) x * f' x = ∫ u in f a..f b, g u`. --/ -theorem integral_comp_mul_deriv {f f' g : ℝ → ℝ} - (h : ∀ x ∈ uIcc a b, has_deriv_at f (f' x) x) - (h' : continuous_on f' (uIcc a b)) (hg : continuous g) : - ∫ x in a..b, (g ∘ f) x * f' x = ∫ x in f a..f b, g x := -integral_comp_mul_deriv' h h' hg.continuous_on - -theorem integral_deriv_comp_mul_deriv' {f f' g g' : ℝ → ℝ} - (hf : continuous_on f [a, b]) - (hff' : ∀ x ∈ Ioo (min a b) (max a b), has_deriv_within_at f (f' x) (Ioi x) x) - (hf' : continuous_on f' [a, b]) - (hg : continuous_on g [f a, f b]) - (hgg' : ∀ x ∈ Ioo (min (f a) (f b)) (max (f a) (f b)), has_deriv_within_at g (g' x) (Ioi x) x) - (hg' : continuous_on g' (f '' [a, b])) : - ∫ x in a..b, (g' ∘ f) x * f' x = (g ∘ f) b - (g ∘ f) a := -by simpa [mul_comm] using integral_deriv_comp_smul_deriv' hf hff' hf' hg hgg' hg' - -theorem integral_deriv_comp_mul_deriv {f f' g g' : ℝ → ℝ} - (hf : ∀ x ∈ uIcc a b, has_deriv_at f (f' x) x) - (hg : ∀ x ∈ uIcc a b, has_deriv_at g (g' (f x)) (f x)) - (hf' : continuous_on f' (uIcc a b)) (hg' : continuous g') : - ∫ x in a..b, (g' ∘ f) x * f' x = (g ∘ f) b - (g ∘ f) a := -by simpa [mul_comm] using integral_deriv_comp_smul_deriv hf hg hf' hg' - -end mul end interval_integral From e4bc74cbaf429d706cb9140902f7ca6c431e75a4 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sun, 30 Apr 2023 07:58:36 +0000 Subject: [PATCH 0909/1291] feat(data/set/intervals/pi): `f.update i '' Icc a b = Icc (f.update i a) (f.update i b)` (#18892) and corresponding lemmas for `pi.single` --- src/algebra/group/pi.lean | 10 ++ src/data/set/intervals/pi.lean | 161 +++++++++++++++++++++++++++++++++ src/order/lattice.lean | 13 +++ 3 files changed, 184 insertions(+) diff --git a/src/algebra/group/pi.lean b/src/algebra/group/pi.lean index d64049b2ec257..5671481eaa151 100644 --- a/src/algebra/group/pi.lean +++ b/src/algebra/group/pi.lean @@ -314,6 +314,16 @@ This is the `mul_hom` version of `pi.single`. -/ variables {f} +@[to_additive] +lemma pi.mul_single_sup [Π i, semilattice_sup (f i)] [Π i, has_one (f i)] (i : I) (x y : f i) : + pi.mul_single i (x ⊔ y) = pi.mul_single i x ⊔ pi.mul_single i y := +function.update_sup _ _ _ _ + +@[to_additive] +lemma pi.mul_single_inf [Π i, semilattice_inf (f i)] [Π i, has_one (f i)] (i : I) (x y : f i) : + pi.mul_single i (x ⊓ y) = pi.mul_single i x ⊓ pi.mul_single i y := +function.update_inf _ _ _ _ + @[to_additive] lemma pi.mul_single_mul [Π i, mul_one_class $ f i] (i : I) (x y : f i) : mul_single i (x * y) = mul_single i x * mul_single i y := diff --git a/src/data/set/intervals/pi.lean b/src/data/set/intervals/pi.lean index 3197184c38759..4377982aa80cb 100644 --- a/src/data/set/intervals/pi.lean +++ b/src/data/set/intervals/pi.lean @@ -3,6 +3,7 @@ Copyright (c) 2020 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ +import data.pi.algebra import data.set.intervals.basic import data.set.intervals.unordered_interval import data.set.lattice @@ -106,12 +107,172 @@ end end pi_preorder +section pi_partial_order +variables [decidable_eq ι] [Π i, partial_order (α i)] + +lemma image_update_Icc (f : Π i, α i) (i : ι) (a b : α i) : + f.update i '' Icc a b = Icc (f.update i a) (f.update i b) := +begin + ext, + rw ←set.pi_univ_Icc, + refine ⟨_, λ h, ⟨x i, _, _⟩⟩, + { rintro ⟨c, hc, rfl⟩, + simpa [update_le_update_iff] }, + { simpa only [function.update_same] using h i (mem_univ i) }, + { ext j, + obtain rfl | hij := eq_or_ne i j, + { exact function.update_same _ _ _ }, + { simpa only [function.update_noteq hij.symm, le_antisymm_iff] using h j (mem_univ j) } } +end + +lemma image_update_Ico (f : Π i, α i) (i : ι) (a b : α i) : + f.update i '' Ico a b = Ico (f.update i a) (f.update i b) := +by rw [←Icc_diff_right, ←Icc_diff_right, image_diff (f.update_injective _), image_singleton, + image_update_Icc] + +lemma image_update_Ioc (f : Π i, α i) (i : ι) (a b : α i) : + f.update i '' Ioc a b = Ioc (f.update i a) (f.update i b) := +by rw [←Icc_diff_left, ←Icc_diff_left, image_diff (f.update_injective _), image_singleton, + image_update_Icc] + +lemma image_update_Ioo (f : Π i, α i) (i : ι) (a b : α i) : + f.update i '' Ioo a b = Ioo (f.update i a) (f.update i b) := +by rw [←Ico_diff_left, ←Ico_diff_left, image_diff (f.update_injective _), image_singleton, + image_update_Ico] + +lemma image_update_Icc_left (f : Π i, α i) (i : ι) (a : α i) : + f.update i '' Icc a (f i) = Icc (f.update i a) f := +by simpa using image_update_Icc f i a (f i) + +lemma image_update_Ico_left (f : Π i, α i) (i : ι) (a : α i) : + f.update i '' Ico a (f i) = Ico (f.update i a) f := +by simpa using image_update_Ico f i a (f i) + +lemma image_update_Ioc_left (f : Π i, α i) (i : ι) (a : α i) : + f.update i '' Ioc a (f i) = Ioc (f.update i a) f := +by simpa using image_update_Ioc f i a (f i) + +lemma image_update_Ioo_left (f : Π i, α i) (i : ι) (a : α i) : + f.update i '' Ioo a (f i) = Ioo (f.update i a) f := +by simpa using image_update_Ioo f i a (f i) + +lemma image_update_Icc_right (f : Π i, α i) (i : ι) (b : α i) : + f.update i '' Icc (f i) b = Icc f (f.update i b) := +by simpa using image_update_Icc f i (f i) b + +lemma image_update_Ico_right (f : Π i, α i) (i : ι) (b : α i) : + f.update i '' Ico (f i) b = Ico f (f.update i b) := +by simpa using image_update_Ico f i (f i) b + +lemma image_update_Ioc_right (f : Π i, α i) (i : ι) (b : α i) : + f.update i '' Ioc (f i) b = Ioc f (f.update i b) := +by simpa using image_update_Ioc f i (f i) b + +lemma image_update_Ioo_right (f : Π i, α i) (i : ι) (b : α i) : + f.update i '' Ioo (f i) b = Ioo f (f.update i b) := +by simpa using image_update_Ioo f i (f i) b + +variables [Π i, has_one (α i)] + +@[to_additive] +lemma image_mul_single_Icc (i : ι) (a b : α i) : + pi.mul_single i '' Icc a b = Icc (pi.mul_single i a) (pi.mul_single i b) := +image_update_Icc _ _ _ _ + +@[to_additive] +lemma image_mul_single_Ico (i : ι) (a b : α i) : + pi.mul_single i '' Ico a b = Ico (pi.mul_single i a) (pi.mul_single i b) := +image_update_Ico _ _ _ _ + +@[to_additive] +lemma image_mul_single_Ioc (i : ι) (a b : α i) : + pi.mul_single i '' Ioc a b = Ioc (pi.mul_single i a) (pi.mul_single i b) := +image_update_Ioc _ _ _ _ + +@[to_additive] +lemma image_mul_single_Ioo (i : ι) (a b : α i) : + pi.mul_single i '' Ioo a b = Ioo (pi.mul_single i a) (pi.mul_single i b) := +image_update_Ioo _ _ _ _ + +@[to_additive] +lemma image_mul_single_Icc_left (i : ι) (a : α i) : + pi.mul_single i '' Icc a 1 = Icc (pi.mul_single i a) 1 := +image_update_Icc_left _ _ _ + +@[to_additive] +lemma image_mul_single_Ico_left (i : ι) (a : α i) : + pi.mul_single i '' Ico a 1 = Ico (pi.mul_single i a) 1 := +image_update_Ico_left _ _ _ + +@[to_additive] +lemma image_mul_single_Ioc_left (i : ι) (a : α i) : + pi.mul_single i '' Ioc a 1 = Ioc (pi.mul_single i a) 1 := +image_update_Ioc_left _ _ _ + +@[to_additive] +lemma image_mul_single_Ioo_left (i : ι) (a : α i) : + pi.mul_single i '' Ioo a 1 = Ioo (pi.mul_single i a) 1 := +image_update_Ioo_left _ _ _ + +@[to_additive] +lemma image_mul_single_Icc_right (i : ι) (b : α i) : + pi.mul_single i '' Icc 1 b = Icc 1 (pi.mul_single i b) := +image_update_Icc_right _ _ _ + +@[to_additive] +lemma image_mul_single_Ico_right (i : ι) (b : α i) : + pi.mul_single i '' Ico 1 b = Ico 1 (pi.mul_single i b) := +image_update_Ico_right _ _ _ + +@[to_additive] +lemma image_mul_single_Ioc_right (i : ι) (b : α i) : + pi.mul_single i '' Ioc 1 b = Ioc 1 (pi.mul_single i b) := +image_update_Ioc_right _ _ _ + +@[to_additive] +lemma image_mul_single_Ioo_right (i : ι) (b : α i) : + pi.mul_single i '' Ioo 1 b = Ioo 1 (pi.mul_single i b) := +image_update_Ioo_right _ _ _ + +end pi_partial_order + section pi_lattice variables [Π i, lattice (α i)] @[simp] lemma pi_univ_uIcc (a b : Π i, α i) : pi univ (λ i, uIcc (a i) (b i)) = uIcc a b := pi_univ_Icc _ _ +variables [decidable_eq ι] + +lemma image_update_uIcc (f : Π i, α i) (i : ι) (a b : α i) : + f.update i '' uIcc a b = uIcc (f.update i a) (f.update i b) := +(image_update_Icc _ _ _ _).trans $ by simp_rw [uIcc, f.update_sup, f.update_inf] + +lemma image_update_uIcc_left (f : Π i, α i) (i : ι) (a : α i) : + f.update i '' uIcc a (f i) = uIcc (f.update i a) f := +by simpa using image_update_uIcc f i a (f i) + +lemma image_update_uIcc_right (f : Π i, α i) (i : ι) (b : α i) : + f.update i '' uIcc (f i) b = uIcc f (f.update i b) := +by simpa using image_update_uIcc f i (f i) b + +variables [Π i, has_one (α i)] + +@[to_additive] +lemma image_mul_single_uIcc (i : ι) (a b : α i) : + pi.mul_single i '' uIcc a b = uIcc (pi.mul_single i a) (pi.mul_single i b) := +image_update_uIcc _ _ _ _ + +@[to_additive] +lemma image_mul_single_uIcc_left (i : ι) (a : α i) : + pi.mul_single i '' uIcc a 1 = uIcc (pi.mul_single i a) 1 := +image_update_uIcc_left _ _ _ + +@[to_additive] +lemma image_mul_single_uIcc_right (i : ι) (b : α i) : + pi.mul_single i '' uIcc 1 b = uIcc 1 (pi.mul_single i b) := +image_update_uIcc_right _ _ _ + end pi_lattice variables [decidable_eq ι] [Π i, linear_order (α i)] diff --git a/src/order/lattice.lean b/src/order/lattice.lean index df7d86f946fb9..e3877be97db81 100644 --- a/src/order/lattice.lean +++ b/src/order/lattice.lean @@ -777,6 +777,19 @@ by refine_struct { .. pi.lattice }; tactic.pi_instance_derive_field end pi +namespace function +variables {ι : Type*} {π : ι → Type*} [decidable_eq ι] + +lemma update_sup [Π i, semilattice_sup (π i)] (f : Π i, π i) (i : ι) (a b : π i) : + f.update i (a ⊔ b) = f.update i a ⊔ f.update i b := +funext $ λ j, by obtain rfl | hji := eq_or_ne j i; simp [update_noteq, *] + +lemma update_inf [Π i, semilattice_inf (π i)] (f : Π i, π i) (i : ι) (a b : π i) : + f.update i (a ⊓ b) = f.update i a ⊓ f.update i b := +funext $ λ j, by obtain rfl | hji := eq_or_ne j i; simp [update_noteq, *] + +end function + /-! ### Monotone functions and lattices -/ From 4367b192b58a665b6f18773f73eb492eb4df7990 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sun, 30 Apr 2023 12:24:59 +0000 Subject: [PATCH 0910/1291] feat(data/set/intervals/basic): generalize from `linear_order` to `preorder` (#18876) The lemma statements are unchanged, but the proofs are rewritten --- src/data/set/intervals/basic.lean | 41 ++++++++++--------------------- 1 file changed, 13 insertions(+), 28 deletions(-) diff --git a/src/data/set/intervals/basic.lean b/src/data/set/intervals/basic.lean index 290b00e84d0fb..1c4c0e2cc4b6d 100644 --- a/src/data/set/intervals/basic.lean +++ b/src/data/set/intervals/basic.lean @@ -407,6 +407,19 @@ lemma _root_.is_min.Iio_eq (h : is_min a) : Iio a = ∅ := eq_empty_of_subset_em lemma Iic_inter_Ioc_of_le (h : a ≤ c) : Iic a ∩ Ioc b c = Ioc b a := ext $ λ x, ⟨λ H, ⟨H.2.1, H.1⟩, λ H, ⟨H.2, H.1, H.2.trans h⟩⟩ +lemma not_mem_Icc_of_lt (ha : c < a) : c ∉ Icc a b := λ h, ha.not_le h.1 +lemma not_mem_Icc_of_gt (hb : b < c) : c ∉ Icc a b := λ h, hb.not_le h.2 +lemma not_mem_Ico_of_lt (ha : c < a) : c ∉ Ico a b := λ h, ha.not_le h.1 +lemma not_mem_Ioc_of_gt (hb : b < c) : c ∉ Ioc a b := λ h, hb.not_le h.2 + +@[simp] lemma not_mem_Ioi_self : a ∉ Ioi a := lt_irrefl _ +@[simp] lemma not_mem_Iio_self : b ∉ Iio b := lt_irrefl _ + +lemma not_mem_Ioc_of_le (ha : c ≤ a) : c ∉ Ioc a b := λ h, lt_irrefl _ $ h.1.trans_le ha +lemma not_mem_Ico_of_ge (hb : b ≤ c) : c ∉ Ico a b := λ h, lt_irrefl _ $ h.2.trans_le hb +lemma not_mem_Ioo_of_le (ha : c ≤ a) : c ∉ Ioo a b := λ h, lt_irrefl _ $ h.1.trans_le ha +lemma not_mem_Ioo_of_ge (hb : b ≤ c) : c ∉ Ioo a b := λ h, lt_irrefl _ $ h.2.trans_le hb + end preorder section partial_order @@ -594,38 +607,10 @@ lemma not_mem_Ici : c ∉ Ici a ↔ c < a := not_le lemma not_mem_Iic : c ∉ Iic b ↔ b < c := not_le -lemma not_mem_Icc_of_lt (ha : c < a) : c ∉ Icc a b := -not_mem_subset Icc_subset_Ici_self $ not_mem_Ici.mpr ha - -lemma not_mem_Icc_of_gt (hb : b < c) : c ∉ Icc a b := -not_mem_subset Icc_subset_Iic_self $ not_mem_Iic.mpr hb - -lemma not_mem_Ico_of_lt (ha : c < a) : c ∉ Ico a b := -not_mem_subset Ico_subset_Ici_self $ not_mem_Ici.mpr ha - -lemma not_mem_Ioc_of_gt (hb : b < c) : c ∉ Ioc a b := -not_mem_subset Ioc_subset_Iic_self $ not_mem_Iic.mpr hb - lemma not_mem_Ioi : c ∉ Ioi a ↔ c ≤ a := not_lt lemma not_mem_Iio : c ∉ Iio b ↔ b ≤ c := not_lt -@[simp] lemma not_mem_Ioi_self : a ∉ Ioi a := lt_irrefl _ - -@[simp] lemma not_mem_Iio_self : b ∉ Iio b := lt_irrefl _ - -lemma not_mem_Ioc_of_le (ha : c ≤ a) : c ∉ Ioc a b := -not_mem_subset Ioc_subset_Ioi_self $ not_mem_Ioi.mpr ha - -lemma not_mem_Ico_of_ge (hb : b ≤ c) : c ∉ Ico a b := -not_mem_subset Ico_subset_Iio_self $ not_mem_Iio.mpr hb - -lemma not_mem_Ioo_of_le (ha : c ≤ a) : c ∉ Ioo a b := -not_mem_subset Ioo_subset_Ioi_self $ not_mem_Ioi.mpr ha - -lemma not_mem_Ioo_of_ge (hb : b ≤ c) : c ∉ Ioo a b := -not_mem_subset Ioo_subset_Iio_self $ not_mem_Iio.mpr hb - @[simp] lemma compl_Iic : (Iic a)ᶜ = Ioi a := ext $ λ _, not_le @[simp] lemma compl_Ici : (Ici a)ᶜ = Iio a := ext $ λ _, not_le @[simp] lemma compl_Iio : (Iio a)ᶜ = Ici a := ext $ λ _, not_lt From b88d81c84530450a8989e918608e5960f015e6c8 Mon Sep 17 00:00:00 2001 From: Peiran Wu <15968905+wupr@users.noreply.github.com> Date: Sun, 30 Apr 2023 15:29:39 +0000 Subject: [PATCH 0911/1291] feat(ring_theory/ideal/quotient_operations): add double quotient of an algebra (#18452) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This adds definitions and lemmas on the double quotient `A / (I ⊔ J)` of a commutative algebra `A` by ideals `I` and `J`, analogous to existing definitions and lemmas on the double quotient of a commutative ring. Co-authored-by: Eric Wieser --- .../ideal/quotient_operations.lean | 160 ++++++++++++++++++ 1 file changed, 160 insertions(+) diff --git a/src/ring_theory/ideal/quotient_operations.lean b/src/ring_theory/ideal/quotient_operations.lean index 8eee8b5ab2641..106c861b9cc01 100644 --- a/src/ring_theory/ideal/quotient_operations.lean +++ b/src/ring_theory/ideal/quotient_operations.lean @@ -517,4 +517,164 @@ lemma quot_quot_equiv_comm_algebra_map (x : R) : end algebra +section algebra_quotient + +variables (R) {A : Type*} [comm_semiring R] [comm_ring A] [algebra R A] +variables (I J : ideal A) + +/-- The natural algebra homomorphism `A / I → A / (I ⊔ J)`. -/ +def quot_left_to_quot_supₐ : A ⧸ I →ₐ[R] A ⧸ (I ⊔ J) := +alg_hom.mk (quot_left_to_quot_sup I J) rfl (map_mul _) rfl (map_add _) (λ _, rfl) + +@[simp] +lemma quot_left_to_quot_supₐ_to_ring_hom : + (quot_left_to_quot_supₐ R I J).to_ring_hom = quot_left_to_quot_sup I J := +rfl + +@[simp] +lemma coe_quot_left_to_quot_supₐ : + ⇑(quot_left_to_quot_supₐ R I J) = quot_left_to_quot_sup I J := +rfl + +/-- The algebra homomorphism `(A / I) / J' -> A / (I ⊔ J) induced by `quot_left_to_quot_sup`, + where `J'` is the projection of `J` in `A / I`. -/ +def quot_quot_to_quot_supₐ : (A ⧸ I) ⧸ J.map (I^.quotient.mkₐ R) →ₐ[R] A ⧸ I ⊔ J := +alg_hom.mk (quot_quot_to_quot_sup I J) rfl (map_mul _) rfl (map_add _) (λ _, rfl) + +@[simp] +lemma quot_quot_to_quot_supₐ_to_ring_hom : + (quot_quot_to_quot_supₐ R I J).to_ring_hom = quot_quot_to_quot_sup I J := +rfl + +@[simp] +lemma coe_quot_quot_to_quot_supₐ : + ⇑(quot_quot_to_quot_supₐ R I J) = quot_quot_to_quot_sup I J := +rfl + +/-- The composition of the algebra homomorphisms `A → (A / I)` and `(A / I) → (A / I) / J'`, + where `J'` is the projection `J` in `A / I`. -/ +def quot_quot_mkₐ : A →ₐ[R] ((A ⧸ I) ⧸ J.map (I^.quotient.mkₐ R)) := +alg_hom.mk (quot_quot_mk I J) rfl (map_mul _) rfl (map_add _) (λ _, rfl) + +@[simp] +lemma quot_quot_mkₐ_to_ring_hom : + (quot_quot_mkₐ R I J).to_ring_hom = quot_quot_mk I J := +rfl + +@[simp] +lemma coe_quot_quot_mkₐ : + ⇑(quot_quot_mkₐ R I J) = quot_quot_mk I J := +rfl + +/-- The injective algebra homomorphism `A / (I ⊔ J) → (A / I) / J'`induced by `quot_quot_mk`, + where `J'` is the projection `J` in `A / I`. -/ +def lift_sup_quot_quot_mkₐ (I J : ideal A) : + A ⧸ (I ⊔ J) →ₐ[R] (A ⧸ I) ⧸ J.map (I^.quotient.mkₐ R) := +alg_hom.mk (lift_sup_quot_quot_mk I J) rfl (map_mul _) rfl (map_add _) (λ _, rfl) + +@[simp] +lemma lift_sup_quot_quot_mkₐ_to_ring_hom : + (lift_sup_quot_quot_mkₐ R I J).to_ring_hom = lift_sup_quot_quot_mk I J := +rfl + +@[simp] +lemma coe_lift_sup_quot_quot_mkₐ : + ⇑(lift_sup_quot_quot_mkₐ R I J) = lift_sup_quot_quot_mk I J := +rfl + +/-- `quot_quot_to_quot_add` and `lift_sup_quot_quot_mk` are inverse isomorphisms. In the case where + `I ≤ J`, this is the Third Isomorphism Theorem (see `quot_quot_equiv_quot_of_le`). -/ +def quot_quot_equiv_quot_supₐ : ((A ⧸ I) ⧸ J.map (I^.quotient.mkₐ R)) ≃ₐ[R] A ⧸ I ⊔ J := +@alg_equiv.of_ring_equiv R _ _ _ _ _ _ _ (quot_quot_equiv_quot_sup I J) (λ _, rfl) + +@[simp] +lemma quot_quot_equiv_quot_supₐ_to_ring_equiv : + (quot_quot_equiv_quot_supₐ R I J).to_ring_equiv = quot_quot_equiv_quot_sup I J := +rfl + +@[simp] +lemma coe_quot_quot_equiv_quot_supₐ : + ⇑(quot_quot_equiv_quot_supₐ R I J) = quot_quot_equiv_quot_sup I J := +rfl + +@[simp] +lemma quot_quot_equiv_quot_supₐ_symm_to_ring_equiv: + (quot_quot_equiv_quot_supₐ R I J).symm.to_ring_equiv = (quot_quot_equiv_quot_sup I J).symm := +rfl + +@[simp] +lemma coe_quot_quot_equiv_quot_supₐ_symm: + ⇑(quot_quot_equiv_quot_supₐ R I J).symm = (quot_quot_equiv_quot_sup I J).symm := +rfl + +/-- The natural algebra isomorphism `(A / I) / J' → (A / J) / I'`, + where `J'` (resp. `I'`) is the projection of `J` in `A / I` (resp. `I` in `A / J`). -/ +def quot_quot_equiv_commₐ : + ((A ⧸ I) ⧸ J.map (I^.quotient.mkₐ R)) ≃ₐ[R] + ((A ⧸ J) ⧸ I.map (J^.quotient.mkₐ R)) := +@alg_equiv.of_ring_equiv R _ _ _ _ _ _ _ (quot_quot_equiv_comm I J) (λ _, rfl) + +@[simp] +lemma quot_quot_equiv_commₐ_to_ring_equiv : + (quot_quot_equiv_commₐ R I J).to_ring_equiv = quot_quot_equiv_comm I J := +rfl + +@[simp] +lemma coe_quot_quot_equiv_commₐ : + ⇑(quot_quot_equiv_commₐ R I J) = quot_quot_equiv_comm I J := +rfl + +@[simp] +lemma quot_quot_equiv_comm_symmₐ : + (quot_quot_equiv_commₐ R I J).symm = quot_quot_equiv_commₐ R J I := +-- TODO: should be `rfl` but times out +alg_equiv.ext $ λ x, (fun_like.congr_fun (quot_quot_equiv_comm_symm I J) x : _) + +@[simp] +lemma quot_quot_equiv_comm_comp_quot_quot_mkₐ : + alg_hom.comp ↑(quot_quot_equiv_commₐ R I J) (quot_quot_mkₐ R I J) = quot_quot_mkₐ R J I := +alg_hom.ext $ quot_quot_equiv_comm_quot_quot_mk I J + +variables {I J} + +/-- The **third isomoprhism theorem** for rings. See `quot_quot_equiv_quot_sup` for version + that does not assume an inclusion of ideals. -/ +def quot_quot_equiv_quot_of_leₐ (h : I ≤ J) : + ((A ⧸ I) ⧸ (J.map (I^.quotient.mkₐ R))) ≃ₐ[R] A ⧸ J := +@alg_equiv.of_ring_equiv R _ _ _ _ _ _ _ (quot_quot_equiv_quot_of_le h) (λ _, rfl) + +@[simp] +lemma quot_quot_equiv_quot_of_leₐ_to_ring_equiv (h : I ≤ J) : + (quot_quot_equiv_quot_of_leₐ R h).to_ring_equiv = quot_quot_equiv_quot_of_le h := +rfl + +@[simp] +lemma coe_quot_quot_equiv_quot_of_leₐ (h : I ≤ J) : + ⇑(quot_quot_equiv_quot_of_leₐ R h) = quot_quot_equiv_quot_of_le h := +rfl + +@[simp] +lemma quot_quot_equiv_quot_of_leₐ_symm_to_ring_equiv (h : I ≤ J) : + (quot_quot_equiv_quot_of_leₐ R h).symm.to_ring_equiv = (quot_quot_equiv_quot_of_le h).symm := +rfl + +@[simp] +lemma coe_quot_quot_equiv_quot_of_leₐ_symm (h : I ≤ J) : + ⇑(quot_quot_equiv_quot_of_leₐ R h).symm = (quot_quot_equiv_quot_of_le h).symm := +rfl + +@[simp] +lemma quot_quot_equiv_quot_of_le_comp_quot_quot_mkₐ (h : I ≤ J) : + alg_hom.comp ↑(quot_quot_equiv_quot_of_leₐ R h) (quot_quot_mkₐ R I J) = + J^.quotient.mkₐ R := +rfl + +@[simp] +lemma quot_quot_equiv_quot_of_le_symm_comp_mkₐ (h : I ≤ J) : + alg_hom.comp ↑(quot_quot_equiv_quot_of_leₐ R h).symm (J^.quotient.mkₐ R) = + quot_quot_mkₐ R I J := +rfl + +end algebra_quotient + end double_quot From 9d684a893c52e1d6692a504a118bfccbae04feeb Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sun, 30 Apr 2023 17:11:43 +0000 Subject: [PATCH 0912/1291] chore(data/finset/lattice): Move lemmas around (#18900) Move `map_finset_sup`/`map_finset_inf` from `order.hom.lattice` to `data.finset.lattice`. This breaks a few unqualified downstream uses of `submodule.map_bot`. --- src/data/finset/lattice.lean | 63 ++++++++++++---------- src/linear_algebra/basic.lean | 2 +- src/linear_algebra/finsupp.lean | 5 +- src/linear_algebra/linear_independent.lean | 6 +-- src/order/hom/complete_lattice.lean | 2 +- src/order/hom/lattice.lean | 13 ----- 6 files changed, 44 insertions(+), 47 deletions(-) diff --git a/src/data/finset/lattice.lean b/src/data/finset/lattice.lean index 828b64da31613..285a898550334 100644 --- a/src/data/finset/lattice.lean +++ b/src/data/finset/lattice.lean @@ -8,6 +8,7 @@ import data.finset.option import data.finset.prod import data.multiset.lattice import order.complete_lattice +import order.hom.lattice /-! # Lattice operations on finsets @@ -16,7 +17,7 @@ import order.complete_lattice > Any changes to this file require a corresponding PR to mathlib4. -/ -variables {α β γ ι : Type*} +variables {F α β γ ι : Type*} namespace finset open multiset order_dual @@ -68,6 +69,10 @@ end theorem sup_congr {f g : β → α} (hs : s₁ = s₂) (hfg : ∀ a ∈ s₂, f a = g a) : s₁.sup f = s₂.sup g := by subst hs; exact finset.fold_congr hfg +@[simp] lemma _root_.map_finset_sup [semilattice_sup β] [order_bot β] [sup_bot_hom_class F α β] + (f : F) (s : finset ι) (g : ι → α) : f (s.sup g) = s.sup (f ∘ g) := +finset.cons_induction_on s (map_bot f) $ λ i s _ h, by rw [sup_cons, sup_cons, map_sup, h] + @[simp] protected lemma sup_le_iff {a : α} : s.sup f ≤ a ↔ (∀ b ∈ s, f b ≤ a) := begin apply iff.trans multiset.sup_le, @@ -292,6 +297,10 @@ lemma inf_inf : s.inf (f ⊓ g) = s.inf f ⊓ s.inf g := theorem inf_congr {f g : β → α} (hs : s₁ = s₂) (hfg : ∀ a ∈ s₂, f a = g a) : s₁.inf f = s₂.inf g := by subst hs; exact finset.fold_congr hfg +@[simp] lemma _root_.map_finset_inf [semilattice_inf β] [order_top β] [inf_top_hom_class F α β] + (f : F) (s : finset ι) (g : ι → α) : f (s.inf g) = s.inf (f ∘ g) := +finset.cons_induction_on s (map_top f) $ λ i s _ h, by rw [inf_cons, inf_cons, map_inf, h] + @[simp] lemma inf_bUnion [decidable_eq β] (s : finset γ) (t : γ → finset β) : (s.bUnion t).inf f = s.inf (λ x, (t x).inf f) := @sup_bUnion αᵒᵈ _ _ _ _ _ _ _ _ @@ -337,32 +346,6 @@ lemma inf_product_right (s : finset β) (t : finset γ) (f : β × γ → α) : @[simp] lemma inf_erase_top [decidable_eq α] (s : finset α) : (s.erase ⊤).inf id = s.inf id := @sup_erase_bot αᵒᵈ _ _ _ _ -lemma sup_sdiff_left {α β : Type*} [boolean_algebra α] (s : finset β) (f : β → α) (a : α) : - s.sup (λ b, a \ f b) = a \ s.inf f := -begin - refine finset.cons_induction_on s _ (λ b t _ h, _), - { rw [sup_empty, inf_empty, sdiff_top] }, - { rw [sup_cons, inf_cons, h, sdiff_inf] } -end - -lemma inf_sdiff_left {α β : Type*} [boolean_algebra α] {s : finset β} (hs : s.nonempty) (f : β → α) - (a : α) : - s.inf (λ b, a \ f b) = a \ s.sup f := -begin - induction hs using finset.nonempty.cons_induction with b b t _ _ h, - { rw [sup_singleton, inf_singleton] }, - { rw [sup_cons, inf_cons, h, sdiff_sup] } -end - -lemma inf_sdiff_right {α β : Type*} [boolean_algebra α] {s : finset β} (hs : s.nonempty) (f : β → α) - (a : α) : - s.inf (λ b, f b \ a) = s.inf f \ a := -begin - induction hs using finset.nonempty.cons_induction with b b t _ _ h, - { rw [inf_singleton, inf_singleton] }, - { rw [inf_cons, inf_cons, h, inf_sdiff] } -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) := @@ -449,6 +432,32 @@ lemma inf_sup_distrib_right (s : finset ι) (f : ι → α) (a : α) : end order_top end distrib_lattice +section boolean_algebra +variables [boolean_algebra α] {s : finset ι} + +lemma sup_sdiff_left (s : finset ι) (f : ι → α) (a : α) : s.sup (λ b, a \ f b) = a \ s.inf f := +begin + refine finset.cons_induction_on s _ (λ b t _ h, _), + { rw [sup_empty, inf_empty, sdiff_top] }, + { rw [sup_cons, inf_cons, h, sdiff_inf] } +end + +lemma inf_sdiff_left (hs : s.nonempty) (f : ι → α) (a : α) : s.inf (λ b, a \ f b) = a \ s.sup f := +begin + induction hs using finset.nonempty.cons_induction with b b t _ _ h, + { rw [sup_singleton, inf_singleton] }, + { rw [sup_cons, inf_cons, h, sdiff_sup] } +end + +lemma inf_sdiff_right (hs : s.nonempty) (f : ι → α) (a : α) : s.inf (λ b, f b \ a) = s.inf f \ a := +begin + induction hs using finset.nonempty.cons_induction with b b t _ _ h, + { rw [inf_singleton, inf_singleton] }, + { rw [inf_cons, inf_cons, h, inf_sdiff] } +end + +end boolean_algebra + section linear_order variables [linear_order α] diff --git a/src/linear_algebra/basic.lean b/src/linear_algebra/basic.lean index 4fe744590ba1e..75b5ff14f46e1 100644 --- a/src/linear_algebra/basic.lean +++ b/src/linear_algebra/basic.lean @@ -1133,7 +1133,7 @@ omit sc lemma ker_cod_restrict {τ₂₁ : R₂ →+* R} (p : submodule R M) (f : M₂ →ₛₗ[τ₂₁] M) (hf) : ker (cod_restrict p f hf) = ker f := -by rw [ker, comap_cod_restrict, map_bot]; refl +by rw [ker, comap_cod_restrict, submodule.map_bot]; refl lemma range_cod_restrict {τ₂₁ : R₂ →+* R} [ring_hom_surjective τ₂₁] (p : submodule R M) (f : M₂ →ₛₗ[τ₂₁] M) (hf) : diff --git a/src/linear_algebra/finsupp.lean b/src/linear_algebra/finsupp.lean index 0b05854234f17..e97cbe551801f 100644 --- a/src/linear_algebra/finsupp.lean +++ b/src/linear_algebra/finsupp.lean @@ -260,7 +260,7 @@ begin haveI := classical.dec_pred (λ x, x ∈ (⋃ i, s i)), suffices : ((submodule.subtype _).comp (restrict_dom M R (⋃ i, s i))).range ≤ ⨆ i, supported M R (s i), - { rwa [linear_map.range_comp, range_restrict_dom, map_top, range_subtype] at this }, + { rwa [linear_map.range_comp, range_restrict_dom, submodule.map_top, range_subtype] at this }, rw [range_le_iff_comap, eq_top_iff], rintro l ⟨⟩, apply finsupp.induction l, { exact zero_mem _ }, @@ -633,7 +633,8 @@ variables {α} {M} {v} theorem total_on_range (s : set α) : (finsupp.total_on α M R v s).range = ⊤ := begin rw [finsupp.total_on, linear_map.range_eq_map, linear_map.map_cod_restrict, - ← linear_map.range_le_iff_comap, range_subtype, map_top, linear_map.range_comp, range_subtype], + ← linear_map.range_le_iff_comap, range_subtype, submodule.map_top, linear_map.range_comp, + range_subtype], exact (span_image_eq_map_total _ _).le end diff --git a/src/linear_algebra/linear_independent.lean b/src/linear_algebra/linear_independent.lean index 6a7e065316d26..7242efa4e2256 100644 --- a/src/linear_algebra/linear_independent.lean +++ b/src/linear_algebra/linear_independent.lean @@ -365,8 +365,8 @@ by apply @linear_independent_comp_subtype_disjoint _ _ _ id theorem linear_independent_iff_total_on {s : set M} : linear_independent R (λ x, x : s → M) ↔ (finsupp.total_on M M R id s).ker = ⊥ := -by rw [finsupp.total_on, linear_map.ker, linear_map.comap_cod_restrict, map_bot, comap_bot, - linear_map.ker_comp, linear_independent_subtype_disjoint, disjoint_iff_inf_le, +by rw [finsupp.total_on, linear_map.ker, linear_map.comap_cod_restrict, submodule.map_bot, + comap_bot, linear_map.ker_comp, linear_independent_subtype_disjoint, disjoint_iff_inf_le, ← map_comap_subtype, map_le_iff_le_comap, comap_bot, ker_subtype, le_bot_iff] lemma linear_independent.restrict_of_comp_subtype {s : set ι} @@ -710,7 +710,7 @@ begin { rw [← linear_map.ker_eq_bot, linear_map.ker_cod_restrict], apply hv }, { rw [← linear_map.range_eq_top, linear_map.range_eq_map, linear_map.map_cod_restrict, - ← linear_map.range_le_iff_comap, range_subtype, map_top], + ← linear_map.range_le_iff_comap, range_subtype, submodule.map_top], rw finsupp.range_total, exact le_rfl }, { intro l, diff --git a/src/order/hom/complete_lattice.lean b/src/order/hom/complete_lattice.lean index 354e1d03d6106..fb0163c8fe28f 100644 --- a/src/order/hom/complete_lattice.lean +++ b/src/order/hom/complete_lattice.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ -import order.complete_lattice +import data.set.lattice import order.hom.lattice /-! diff --git a/src/order/hom/lattice.lean b/src/order/hom/lattice.lean index 7bda628e94690..f9fd30def0bdf 100644 --- a/src/order/hom/lattice.lean +++ b/src/order/hom/lattice.lean @@ -3,7 +3,6 @@ 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.lattice import order.hom.bounded import order.symm_diff @@ -216,18 +215,6 @@ instance order_iso_class.to_bounded_lattice_hom_class [lattice α] [lattice β] bounded_lattice_hom_class F α β := { ..order_iso_class.to_lattice_hom_class, ..order_iso_class.to_bounded_order_hom_class } -@[simp] lemma map_finset_sup [semilattice_sup α] [order_bot α] [semilattice_sup β] [order_bot β] - [sup_bot_hom_class F α β] (f : F) (s : finset ι) (g : ι → α) : - f (s.sup g) = s.sup (f ∘ g) := -finset.cons_induction_on s (map_bot f) $ λ i s _ h, - by rw [finset.sup_cons, finset.sup_cons, map_sup, h] - -@[simp] lemma map_finset_inf [semilattice_inf α] [order_top α] [semilattice_inf β] [order_top β] - [inf_top_hom_class F α β] (f : F) (s : finset ι) (g : ι → α) : - f (s.inf g) = s.inf (f ∘ g) := -finset.cons_induction_on s (map_top f) $ λ i s _ h, - by rw [finset.inf_cons, finset.inf_cons, map_inf, h] - section bounded_lattice variables [lattice α] [bounded_order α] [lattice β] [bounded_order β] [bounded_lattice_hom_class F α β] (f : F) {a b : α} From 9425b6f8220e53b059f5a4904786c3c4b50fc057 Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Sun, 30 Apr 2023 20:09:09 +0000 Subject: [PATCH 0913/1291] feat(analysis/fourier): Riemann-Lebesgue lemma (#17719) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR adds a proof that for any function `f` on a finite-dimensional real vector space `V`, the integral `∫ (v : V), exp (-2 * π * w v * I) • f x` tends to 0 as `w` goes to infinity (wrt the cocompact filter on the dual space `W` of `V`). Co-authored-by: sgouezel --- docs/undergrad.yaml | 2 +- src/analysis/fourier/fourier_transform.lean | 54 +-- .../fourier/riemann_lebesgue_lemma.lean | 379 ++++++++++++------ .../inner_product_space/euclidean_dist.lean | 25 +- .../normed_space/finite_dimension.lean | 54 --- .../algebra/module/finite_dimension.lean | 63 ++- 6 files changed, 368 insertions(+), 209 deletions(-) diff --git a/docs/undergrad.yaml b/docs/undergrad.yaml index b9fe40e3e405a..2b45ce944747e 100644 --- a/docs/undergrad.yaml +++ b/docs/undergrad.yaml @@ -527,7 +527,7 @@ Measures and integral calculus: regularization by convolution: 'has_compact_support.cont_diff_convolution_left' Fourier analysis: Fourier series of locally integrable periodic real-valued functions: 'fourier_basis' - Riemann-Lebesgue lemma: '' + Riemann-Lebesgue lemma: 'tendsto_integral_exp_smul_cocompact' convolution product of periodic functions: '' Dirichlet theorem: '' Fejer theorem: '' diff --git a/src/analysis/fourier/fourier_transform.lean b/src/analysis/fourier/fourier_transform.lean index 75287e8e09d99..cea8feb24a888 100644 --- a/src/analysis/fourier/fourier_transform.lean +++ b/src/analysis/fourier/fourier_transform.lean @@ -86,11 +86,10 @@ begin end /-- The uniform norm of the Fourier integral of `f` is bounded by the `L¹` norm of `f`. -/ -lemma fourier_integral_norm_le (e : (multiplicative 𝕜) →* 𝕊) {μ : measure V} - (L : V →ₗ[𝕜] W →ₗ[𝕜] 𝕜) {f : V → E} (hf : integrable f μ) (w : W) : - ‖fourier_integral e μ L f w‖ ≤ ‖hf.to_L1 f‖ := +lemma norm_fourier_integral_le_integral_norm (e : (multiplicative 𝕜) →* 𝕊) (μ : measure V) + (L : V →ₗ[𝕜] W →ₗ[𝕜] 𝕜) (f : V → E) (w : W) : + ‖fourier_integral e μ L f w‖ ≤ ∫ (v : V), ‖f v‖ ∂μ := begin - rw L1.norm_of_fun_eq_integral_norm, refine (norm_integral_le_integral_norm _).trans (le_of_eq _), simp_rw [norm_smul, complex.norm_eq_abs, abs_coe_circle, one_mul], end @@ -126,20 +125,29 @@ section continuous variables [topological_space 𝕜] [topological_ring 𝕜] [topological_space V] [borel_space V] [topological_space W] {e : (multiplicative 𝕜) →* 𝕊} {μ : measure V} {L : V →ₗ[𝕜] W →ₗ[𝕜] 𝕜} -/-- If `f` is integrable, then the Fourier integral is convergent for all `w`. -/ -lemma fourier_integral_convergent - (he : continuous e) (hL : continuous (λ p : V × W, L p.1 p.2)) - {f : V → E} (hf : integrable f μ) (w : W) : - integrable (λ (v : V), (e [-L v w]) • f v) μ := +/-- For any `w`, the Fourier integral is convergent iff `f` is integrable. -/ +lemma fourier_integral_convergent_iff (he : continuous e) (hL : continuous (λ p : V × W, L p.1 p.2)) + {f : V → E} (w : W) : + integrable f μ ↔ integrable (λ (v : V), (e [-L v w]) • f v) μ := begin - rw continuous_induced_rng at he, - have c : continuous (λ v, e[-L v w]), - { refine he.comp (continuous_of_add.comp (continuous.neg _)), - exact hL.comp (continuous_prod_mk.mpr ⟨continuous_id, continuous_const⟩) }, - rw ←integrable_norm_iff (c.ae_strongly_measurable.smul hf.1), - convert hf.norm, + -- first prove one-way implication + have aux : ∀ {g : V → E} (hg : integrable g μ) (x : W), + integrable (λ (v : V), (e [-L v x]) • g v) μ, + { intros g hg x, + have c : continuous (λ v, e[-L v x]), + { refine (continuous_induced_rng.mp he).comp (continuous_of_add.comp (continuous.neg _)), + exact hL.comp (continuous_prod_mk.mpr ⟨continuous_id, continuous_const⟩) }, + rw ←integrable_norm_iff (c.ae_strongly_measurable.smul hg.1), + convert hg.norm, + ext1 v, + rw [norm_smul, complex.norm_eq_abs, abs_coe_circle, one_mul] }, + -- then use it for both directions + refine ⟨λ hf, aux hf w, λ hf, _⟩, + convert aux hf (-w), ext1 v, - rw [norm_smul, complex.norm_eq_abs, abs_coe_circle, one_mul] + rw [←smul_assoc, smul_eq_mul, ←submonoid.coe_mul, ←monoid_hom.map_mul, + ←of_add_add, linear_map.map_neg, neg_neg, ←sub_eq_add_neg, sub_self, of_add_zero, + monoid_hom.map_one, submonoid.coe_one, one_smul], end variables [complete_space E] @@ -153,8 +161,8 @@ begin dsimp only [pi.add_apply, fourier_integral], simp_rw smul_add, rw integral_add, - { exact fourier_integral_convergent he hL hf w }, - { exact fourier_integral_convergent he hL hg w }, + { exact (fourier_integral_convergent_iff he hL w).mp hf }, + { exact (fourier_integral_convergent_iff he hL w).mp hg }, end /-- The Fourier integral of an `L^1` function is a continuous function. -/ @@ -164,7 +172,7 @@ lemma fourier_integral_continuous [topological_space.first_countable_topology W] continuous (fourier_integral e μ L f) := begin apply continuous_of_dominated, - { exact λ w, (fourier_integral_convergent he hL hf w).1 }, + { exact λ w, ((fourier_integral_convergent_iff he hL w).mp hf).1 }, { refine λ w, ae_of_all _ (λ v, _), { exact λ v, ‖f v‖ }, { rw [norm_smul, complex.norm_eq_abs, abs_coe_circle, one_mul] } }, @@ -204,10 +212,10 @@ lemma fourier_integral_smul_const vector_fourier.fourier_integral_smul_const _ _ _ _ _ /-- The uniform norm of the Fourier transform of `f` is bounded by the `L¹` norm of `f`. -/ -lemma fourier_integral_norm_le (e : (multiplicative 𝕜) →* 𝕊) {μ : measure 𝕜} - {f : 𝕜 → E} (hf : integrable f μ) (w : 𝕜) : - ‖fourier_integral e μ f w‖ ≤ ‖hf.to_L1 f‖ := -vector_fourier.fourier_integral_norm_le _ _ _ _ +lemma norm_fourier_integral_le_integral_norm + (e : (multiplicative 𝕜) →* 𝕊) (μ : measure 𝕜) (f : 𝕜 → E) (w : 𝕜) : + ‖fourier_integral e μ f w‖ ≤ ∫ x : 𝕜, ‖f x‖ ∂μ := +vector_fourier.norm_fourier_integral_le_integral_norm _ _ _ _ _ /-- The Fourier transform converts right-translation into scalar multiplication by a phase factor.-/ lemma fourier_integral_comp_add_right [has_measurable_add 𝕜] diff --git a/src/analysis/fourier/riemann_lebesgue_lemma.lean b/src/analysis/fourier/riemann_lebesgue_lemma.lean index de839b44c1f68..5813bab5b2938 100644 --- a/src/analysis/fourier/riemann_lebesgue_lemma.lean +++ b/src/analysis/fourier/riemann_lebesgue_lemma.lean @@ -9,175 +9,312 @@ import measure_theory.integral.integral_eq_improper import measure_theory.group.integration import topology.continuous_function.zero_at_infty import analysis.fourier.fourier_transform +import analysis.inner_product_space.dual +import topology.metric_space.emetric_paracompact +import analysis.inner_product_space.euclidean_dist /-! # The Riemann-Lebesgue Lemma -In this file we prove a weak form of the Riemann-Lebesgue lemma, stating that for any -compactly-supported continuous function `f` on `ℝ` (valued in some complete normed space `E`), the -integral +In this file we prove the Riemann-Lebesgue lemma, for functions on finite-dimensional real vector +spaces `V`: if `f` is a function on `V` (valued in a complete normed space `E`), then the +Fourier transform of `f`, viewed as a function on the dual space of `V`, tends to 0 along the +cocompact filter. Here the Fourier transform is defined by -`∫ (x : ℝ), exp (↑(t * x) * I) • f x` +`λ w : V →L[ℝ] ℝ, ∫ (v : V), exp (↑(2 * π * w v) * I) • f x`. -tends to zero as `t → ∞`. (The actual lemma is that this holds for all `L¹` functions `f`, which -follows from the result proved here together with the fact that continuous, compactly-supported -functions are dense in `L¹(ℝ)`, which will be proved in a future iteration.) +This is true for arbitrary functions, but is only interesting for `L¹` functions (if `f` is not +integrable then the integral is zero for all `w`). This is proved first for continuous +compactly-supported functions on inner-product spaces; then we pass to arbitrary functions using the +density of continuous compactly-supported functions in `L¹` space. Finally we generalise from +inner-product spaces to arbitrary finite-dimensional spaces, by choosing a continuous linear +equivalence to an inner-product space. ## Main results -- `tendsto_integral_mul_exp_at_top_of_continuous_compact_support`: the Riemann-Lebesgue lemma for - continuous compactly-supported functions on `ℝ`. +- `tendsto_integral_exp_inner_smul_cocompact` : for `V` a finite-dimensional real inner product + space and `f : V → E`, the function `λ w : V, ∫ v : V, exp (2 * π * ⟪w, v⟫ * I) • f v` tends to 0 + along `cocompact V`. +- `tendsto_integral_exp_smul_cocompact` : for `V` a finite-dimensional real vector space (endowed + with its unique Hausdorff topological vector space structure), and `W` the dual of `V`, the + function `λ w : W, ∫ v : V, exp (2 * π * w v * I) • f v` tends to along `cocompact W`. +- `real.tendsto_integral_exp_smul_cocompact`: special case of functions on `ℝ`. +- `real.zero_at_infty_fourier_integral` and `real.zero_at_infty_vector_fourier_integral`: + reformulations explicitly using the Fourier integral. -/ +noncomputable theory -open measure_theory filter complex set -open_locale filter topology real ennreal +open measure_theory filter complex set finite_dimensional +open_locale filter topology real ennreal fourier_transform real_inner_product_space nnreal -section continuous_compact_support +variables {E V : Type*} [normed_add_comm_group E] [normed_space ℂ E] {f : V → E} -variables {E : Type*} [normed_add_comm_group E] [normed_space ℂ E] {f : ℝ → E} +local notation `e` := real.fourier_char -/-- The integrand in the Riemann-Lebesgue lemma is integrable. -/ -lemma fourier_integrand_integrable (hf : integrable f) (t : ℝ) : - integrable (λ x:ℝ, exp (↑(t * x) * I) • f x) := +section inner_product_space + +variables [normed_add_comm_group V] [measurable_space V] [borel_space V] + [inner_product_space ℝ V] [finite_dimensional ℝ V] + +/-- The integrand in the Riemann-Lebesgue lemma for `f` is integrable iff `f` is. -/ +lemma fourier_integrand_integrable (w : V) : + integrable f ↔ integrable (λ v : V, e [-⟪v, w⟫] • f v) := begin - rw ←integrable_norm_iff, - simp_rw [norm_smul, norm_exp_of_real_mul_I, one_mul], - exacts [hf.norm, (continuous.ae_strongly_measurable (by continuity)).smul hf.1], + have hL : continuous (λ p : V × V, bilin_form_of_real_inner.to_lin p.1 p.2) := continuous_inner, + rw vector_fourier.fourier_integral_convergent_iff real.continuous_fourier_char hL w, + { simp only [bilin_form.to_lin_apply, bilin_form_of_real_inner_apply] }, + { apply_instance }, end -variable [complete_space E] +variables [complete_space E] -/-- Shifting `f` by `π / t` negates the integral in the Riemann-Lebesgue lemma. -/ -lemma fourier_integral_half_period_translate {t : ℝ} (ht : t ≠ 0) : - ∫ x:ℝ, exp (↑(t * x) * I) • f (x + π / t) = -∫ x:ℝ, exp (↑(t * x) * I) • f x := +local notation `i` := λ w, (1 / (2 * ‖w‖ ^ 2)) • w + +/-- Shifting `f` by `(1 / (2 * ‖w‖ ^ 2)) • w` negates the integral in the Riemann-Lebesgue lemma. -/ +lemma fourier_integral_half_period_translate {w : V} (hw : w ≠ 0) : + ∫ (v : V), e [-⟪v, w⟫] • f (v + i w) = -∫ (v : V), e [-⟪v, w⟫] • f v := begin - have : (λ x:ℝ, exp (↑(t * x) * I) • f (x + π / t)) = - (λ x:ℝ, (λ y:ℝ, -exp (↑(t * y) * I) • f y) (x + π / t)), - { ext1 x, dsimp only, - rw [of_real_mul, of_real_mul, of_real_add, mul_add, add_mul, exp_add, ←neg_mul], - replace ht := complex.of_real_ne_zero.mpr ht, - have : ↑t * ↑(π / t) * I = π * I, by { field_simp, ring }, - rw [this, exp_pi_mul_I], ring_nf, }, + have hiw : ⟪i w, w⟫ = 1 / 2, + { rw [inner_smul_left, inner_self_eq_norm_sq_to_K, is_R_or_C.coe_real_eq_id, id.def, + is_R_or_C.conj_to_real, ←div_div, div_mul_cancel], + rwa [ne.def, sq_eq_zero_iff, norm_eq_zero] }, + have : (λ v : V, e [-⟪v, w⟫] • f (v + i w)) = (λ v : V, (λ x : V, -e[-⟪x, w⟫] • f x) (v + i w)), + { ext1 v, + simp_rw [inner_add_left, hiw, real.fourier_char_apply, neg_add, mul_add, of_real_add, add_mul, + exp_add], + have : 2 * π * -(1 / 2) = -π, by { field_simp, ring }, + rw [this, of_real_neg, neg_mul, exp_neg, exp_pi_mul_I, inv_neg, inv_one, mul_neg_one, neg_neg]}, rw [this, integral_add_right_eq_self], - simp_rw [neg_smul, integral_neg], + simp only [neg_smul, integral_neg], end -/-- Rewrite the Riemann-Lebesgue integral in a form that allows us to use uniform continuity. -/ -lemma fourier_integral_eq_half_sub_half_period_translate - {t : ℝ} (ht : t ≠ 0) (hf : integrable f) : - ∫ x:ℝ, exp (↑(t * x) * I) • f x = - (1 / (2 : ℂ)) • ∫ x:ℝ, exp (↑(t * x) * I) • (f x - f (x + π / t)) := +/-- Rewrite the Fourier integral in a form that allows us to use uniform continuity. -/ +lemma fourier_integral_eq_half_sub_half_period_translate {w : V} (hw : w ≠ 0) (hf : integrable f) : + ∫ v : V, e[-⟪v, w⟫] • f v = (1 / (2 : ℂ)) • ∫ v : V, e[-⟪v, w⟫] • (f v - f (v + i w)) := begin simp_rw [smul_sub], - rw [integral_sub, fourier_integral_half_period_translate ht, sub_eq_add_neg, neg_neg, + rw [integral_sub, fourier_integral_half_period_translate hw, sub_eq_add_neg, neg_neg, ←two_smul ℂ _, ←@smul_assoc _ _ _ _ _ _ (is_scalar_tower.left ℂ), smul_eq_mul], norm_num, - exacts [fourier_integrand_integrable hf t, - fourier_integrand_integrable (hf.comp_add_right (π / t)) t], + exacts [(fourier_integrand_integrable w).mp hf, + (fourier_integrand_integrable w).mp (hf.comp_add_right _)], end /-- Riemann-Lebesgue Lemma for continuous and compactly-supported functions: the integral -`∫ x, exp (t * x * I) • f x` tends to 0 as `t` gets large. -/ -lemma tendsto_integral_mul_exp_at_top_of_continuous_compact_support +`∫ v, exp (-2 * π * ⟪w, v⟫ * I) • f v` tends to 0 wrt `cocompact V`. Note that this is primarily +of interest as a preparatory step for the more general result +`tendsto_integral_exp_inner_smul_cocompact` in which `f` can be arbitrary. -/ +lemma tendsto_integral_exp_inner_smul_cocompact_of_continuous_compact_support (hf1 : continuous f) (hf2 : has_compact_support f) : - tendsto (λ t:ℝ, ∫ x:ℝ, exp (↑(t * x) * I) • f x) at_top (𝓝 0) := + tendsto (λ w : V, ∫ v : V, e[-⟪v, w⟫] • f v) (cocompact V) (𝓝 0) := begin - simp_rw [normed_add_comm_group.tendsto_nhds_zero, eventually_at_top, ge_iff_le], - intros ε hε, - -- Extract an explicit candidate bound on `t` from uniform continuity. - obtain ⟨R, hR1, hR2⟩ := hf2.exists_pos_le_norm, + refine normed_add_comm_group.tendsto_nhds_zero.mpr (λ ε hε, _), + suffices : ∃ (T : ℝ), ∀ (w : V), T ≤ ‖w‖ → ‖∫ (v : V), e[-⟪v, w⟫] • f v‖ < ε, + { simp_rw [←comap_dist_left_at_top_eq_cocompact (0 : V), eventually_comap, eventually_at_top, + dist_eq_norm', sub_zero], + exact let ⟨T, hT⟩ := this in ⟨T, (λ b hb v hv, hT v (hv.symm ▸ hb))⟩ }, + obtain ⟨R, hR_pos, hR_bd⟩ : ∃ (R : ℝ), 0 < R ∧ ∀ (x : V), R ≤ ‖x‖ → f x = 0, + from hf2.exists_pos_le_norm, + let A := {v : V | ‖v‖ ≤ R + 1}, + have mA : measurable_set A, + { suffices : A = metric.closed_ball (0 : V) (R + 1), + by { rw this, exact metric.is_closed_ball.measurable_set }, + simp_rw [A, metric.closed_ball, dist_eq_norm, sub_zero] }, + obtain ⟨B, hB_pos, hB_vol⟩ : ∃ (B : ℝ≥0), 0 < B ∧ volume A ≤ B, + { have hc : is_compact A, by simpa only [metric.closed_ball, dist_eq_norm, sub_zero] + using is_compact_closed_ball (0 : V) _, + let B₀ := volume A, + replace hc : B₀ < ⊤ := hc.measure_lt_top, + refine ⟨B₀.to_nnreal + 1, add_pos_of_nonneg_of_pos B₀.to_nnreal.coe_nonneg one_pos, _⟩, + rw [ennreal.coe_add, ennreal.coe_one, ennreal.coe_to_nnreal hc.ne], + exact le_self_add }, + --* Use uniform continuity to choose δ such that `‖x - y‖ < δ` implies `‖f x - f y‖ < ε / B`. obtain ⟨δ, hδ1, hδ2⟩ := metric.uniform_continuous_iff.mp - (hf2.uniform_continuous_of_continuous hf1) (ε / (1 + 2 * R)) (div_pos hε (by positivity)), - refine ⟨max π (1 + π / δ), λ t ht, _⟩, - have tpos : 0 < t := lt_of_lt_of_le real.pi_pos ((le_max_left _ _).trans ht), - -- Rewrite integral in terms of `f x - f (x + π / t)`. - rw fourier_integral_eq_half_sub_half_period_translate (lt_of_lt_of_le - (lt_max_of_lt_left real.pi_pos) ht).ne' (hf1.integrable_of_has_compact_support hf2), - rw [norm_smul, norm_eq_abs, ←complex.of_real_one, ←of_real_bit0, ←of_real_div, - complex.abs_of_nonneg one_half_pos.le], - have : ε = (1 / 2) * (2 * ε), by { field_simp, ring, }, + (hf2.uniform_continuous_of_continuous hf1) (ε / B) (div_pos hε hB_pos), + refine ⟨1 / 2 + 1 / (2 * δ), λ w hw_bd, _⟩, + have hw_ne : w ≠ 0, + { contrapose! hw_bd, rw [hw_bd, norm_zero], + exact add_pos one_half_pos (one_div_pos.mpr $ mul_pos two_pos hδ1) }, + have hw'_nm : ‖i w‖ = 1 / (2 * ‖w‖), + { rw [norm_smul, norm_div, real.norm_of_nonneg (mul_nonneg two_pos.le $ sq_nonneg _), norm_one, + sq, ←div_div, ←div_div, ←div_div, div_mul_cancel _ (norm_eq_zero.not.mpr hw_ne)] }, + --* Rewrite integral in terms of `f v - f (v + w')`. + rw [fourier_integral_eq_half_sub_half_period_translate hw_ne + (hf1.integrable_of_has_compact_support hf2), norm_smul, norm_eq_abs, ←complex.of_real_one, + ←of_real_bit0, ←of_real_div, complex.abs_of_nonneg one_half_pos.le], + have : ε = (1 / 2) * (2 * ε), by { field_simp, rw mul_comm }, rw [this, mul_lt_mul_left (one_half_pos : (0:ℝ) < 1/2)], - have : ‖∫ (x : ℝ), exp (↑(t * x) * I) • (f x - f (x + π / t))‖ ≤ ∫ (x : ℝ), - ‖exp (↑(t * x) * I) • (f x - f (x + π / t))‖, from norm_integral_le_integral_norm _, - refine lt_of_le_of_lt this _, - simp_rw [norm_smul, norm_exp_of_real_mul_I, one_mul], - -- Show integral can be taken over `[-(R + 1), R] ⊂ ℝ`. - let A := Icc (-(R + 1)) R, - have int_Icc : ∫ (x : ℝ), ‖f x - f (x + π / t)‖ = ∫ x in A, ‖f x - f (x + π / t)‖, - { refine (set_integral_eq_integral_of_forall_compl_eq_zero (λ x hx, _)).symm, - rw [mem_Icc, not_and_distrib, not_le, not_le, lt_neg] at hx, - suffices : (f x = 0 ∧ f (x + π / t) = 0), by { rw [this.1, this.2, sub_zero, norm_zero], }, - have tp : 0 < t := real.pi_pos.trans_le ((le_max_left _ _).trans ht), - refine ⟨hR2 x $ le_abs.mpr _, hR2 _ $ le_abs.mpr _⟩, - { cases hx, - { exact or.inr ((le_add_of_nonneg_right zero_le_one).trans hx.le) }, - { exact or.inl hx.le } }, - { cases hx, - { refine or.inr _, - rw [neg_add, ←sub_eq_add_neg, le_sub_iff_add_le], - refine le_trans (add_le_add_left _ R) hx.le, - exact (div_le_one tp).mpr ((le_max_left _ _).trans ht) }, - { exact or.inl (hx.trans $ lt_add_of_pos_right _ $ div_pos real.pi_pos tp).le } } }, - rw int_Icc, - -- Bound integral using fact that ‖f x - f (x + π / t)‖ is small. - have bdA : ∀ x : ℝ, (x ∈ A) → ‖ ‖f x - f (x + π / t) ‖ ‖ ≤ ε / (1 + 2 * R), + refine lt_of_le_of_lt (norm_integral_le_integral_norm _) _, + simp_rw [norm_smul, norm_eq_abs, abs_coe_circle, one_mul], + --* Show integral can be taken over A only. + have int_A : ∫ (v : V), ‖f v - f (v + i w)‖ = ∫ v in A, ‖f v - f (v + i w)‖, + { refine (set_integral_eq_integral_of_forall_compl_eq_zero (λ v hv, _)).symm, + dsimp only [A] at hv, + simp only [A, mem_set_of_eq, not_le] at hv, + rw [hR_bd v _, hR_bd (v + i w) _, sub_zero, norm_zero], + { rw ←sub_neg_eq_add, + refine le_trans _ (norm_sub_norm_le _ _), + rw [le_sub_iff_add_le, norm_neg], + refine le_trans _ hv.le, + rw [add_le_add_iff_left, hw'_nm, ←div_div], + refine (div_le_one $ norm_pos_iff.mpr hw_ne).mpr _, + refine le_trans (le_add_of_nonneg_right $ one_div_nonneg.mpr $ _) hw_bd, + exact (mul_pos (zero_lt_two' ℝ) hδ1).le }, + { exact ((le_add_iff_nonneg_right _).mpr zero_le_one).trans hv.le } }, + rw int_A, clear int_A, + --* Bound integral using fact that `‖f v - f (v + w')‖` is small. + have bdA : ∀ v : V, (v ∈ A) → ‖ ‖f v - f (v + i w) ‖ ‖ ≤ ε / B, { simp_rw norm_norm, - refine (λ x _, le_of_lt _), simp_rw dist_eq_norm at hδ2, - apply hδ2, - rw [sub_add_cancel', real.norm_eq_abs, abs_neg, abs_of_pos (div_pos real.pi_pos tpos), - div_lt_iff tpos, mul_comm, ←div_lt_iff hδ1], - linarith [(le_max_right π (1 + π / δ)).trans ht] }, - have bdA2 := norm_set_integral_le_of_norm_le_const (measure_Icc_lt_top : volume A < ∞) bdA _, + refine (λ x _, (hδ2 _).le), + rw [sub_add_cancel', norm_neg, hw'_nm, ←div_div, div_lt_iff (norm_pos_iff.mpr hw_ne), + ←div_lt_iff' hδ1, div_div], + refine (lt_add_of_pos_left _ _).trans_le hw_bd, + exact one_half_pos, }, + have bdA2 := norm_set_integral_le_of_norm_le_const (hB_vol.trans_lt ennreal.coe_lt_top) bdA _, swap, { apply continuous.ae_strongly_measurable, exact (continuous_norm.comp $ continuous.sub hf1 $ continuous.comp hf1 $ continuous_id'.add continuous_const) }, - have : ‖ _ ‖ = ∫ (x : ℝ) in A, ‖f x - f (x + π / t)‖ := - real.norm_of_nonneg (set_integral_nonneg measurable_set_Icc (λ x hx, norm_nonneg _)), + have : ‖ _ ‖ = ∫ (v : V) in A, ‖f v - f (v + i w)‖ := + real.norm_of_nonneg (set_integral_nonneg mA (λ x hx, norm_nonneg _)), rw this at bdA2, - refine lt_of_le_of_lt bdA2 _, - rw [real.volume_Icc, (by ring : R - (-(R + 1)) = 1 + 2 * R)], - have hh : 0 < 1 + 2 * R, by positivity, - rw [ennreal.to_real_of_real hh.le, div_mul_cancel _ hh.ne', two_mul], - exact lt_add_of_pos_left _ hε, + refine bdA2.trans_lt _, + rw [div_mul_eq_mul_div, div_lt_iff (nnreal.coe_pos.mpr hB_pos), mul_comm (2 : ℝ), mul_assoc, + mul_lt_mul_left hε], + rw ← ennreal.to_real_le_to_real at hB_vol, + { refine hB_vol.trans_lt _, + rw [(by refl : (↑B : ennreal).to_real = ↑B), two_mul], + exact lt_add_of_pos_left _ hB_pos }, + exacts [(hB_vol.trans_lt ennreal.coe_lt_top).ne, ennreal.coe_lt_top.ne], end -lemma tendsto_integral_mul_exp_at_bot_of_continuous_compact_support - (hf1 : continuous f) (hf2 : has_compact_support f) : - tendsto (λ t:ℝ, ∫ x:ℝ, exp (↑(t * x) * I) • f x) at_bot (𝓝 0) := +variables (f) + +/-- Riemann-Lebesgue lemma for functions on a real inner-product space: the integral +`∫ v, exp (-2 * π * ⟪w, v⟫ * I) • f v` tends to 0 as `w → ∞`. -/ +theorem tendsto_integral_exp_inner_smul_cocompact : + tendsto (λ w : V, ∫ v, e [-⟪v, w⟫] • f v) (cocompact V) (𝓝 0) := begin - have hg2 : has_compact_support (f ∘ has_neg.neg), - by simpa only [neg_one_smul] using hf2.comp_smul (neg_ne_zero.mpr $ one_ne_zero' ℝ), - convert (tendsto_integral_mul_exp_at_top_of_continuous_compact_support (hf1.comp continuous_neg) - hg2).comp tendsto_neg_at_bot_at_top, - ext1 t, - simp_rw [function.comp_app, neg_mul, ←mul_neg], - rw ←integral_neg_eq_self, + by_cases hfi : integrable f, swap, + { convert tendsto_const_nhds, + ext1 w, + apply integral_undef, + rwa ←fourier_integrand_integrable w }, + refine metric.tendsto_nhds.mpr (λ ε hε, _), + obtain ⟨g, hg_supp, hfg, hg_cont, -⟩ := + hfi.exists_has_compact_support_integral_sub_le (div_pos hε two_pos), + refine ((metric.tendsto_nhds.mp + (tendsto_integral_exp_inner_smul_cocompact_of_continuous_compact_support hg_cont hg_supp)) _ + (div_pos hε two_pos)).mp (eventually_of_forall (λ w hI, _)), + rw dist_eq_norm at hI ⊢, + have : ‖(∫ v, e [-⟪v, w⟫] • f v) - (∫ v, e [-⟪v, w⟫] • g v)‖ ≤ ε / 2, + { refine le_trans _ hfg, + simp_rw [←integral_sub ((fourier_integrand_integrable w).mp hfi) + ((fourier_integrand_integrable w).mp (hg_cont.integrable_of_has_compact_support hg_supp)), + ←smul_sub, ←pi.sub_apply], + exact vector_fourier.norm_fourier_integral_le_integral_norm e volume + bilin_form_of_real_inner.to_lin (f - g) w }, + replace := add_lt_add_of_le_of_lt this hI, + rw add_halves at this, + refine ((le_of_eq _).trans (norm_add_le _ _)).trans_lt this, + simp only [sub_zero, sub_add_cancel], end -lemma zero_at_infty_integral_mul_exp_of_continuous_compact_support - (hf1 : continuous f) (hf2 : has_compact_support f) : - tendsto (λ t:ℝ, ∫ x:ℝ, exp (↑(t * x) * I) • f x) (cocompact ℝ) (𝓝 0) := +/-- The Riemann-Lebesgue lemma for functions on `ℝ`. -/ +lemma real.tendsto_integral_exp_smul_cocompact (f : ℝ → E) : + tendsto (λ w : ℝ, ∫ v : ℝ, e [-(v * w)] • f v) (cocompact ℝ) (𝓝 0) := +tendsto_integral_exp_inner_smul_cocompact f + +/-- The Riemann-Lebesgue lemma for functions on `ℝ`, formulated via `real.fourier_integral`. -/ +theorem real.zero_at_infty_fourier_integral (f : ℝ → E) : + tendsto (𝓕 f) (cocompact ℝ) (𝓝 0) := +tendsto_integral_exp_inner_smul_cocompact f + +/-- Riemann-Lebesgue lemma for functions on a finite-dimensional inner-product space, formulated +via dual space. **Do not use** -- it is only a stepping stone to +`tendsto_integral_exp_smul_cocompact` where the inner-product-space structure isn't required. -/ +lemma tendsto_integral_exp_smul_cocompact_of_inner_product (μ : measure V) [μ.is_add_haar_measure] : + tendsto (λ w : V →L[ℝ] ℝ, ∫ v, e[-w v] • f v ∂μ) (cocompact (V →L[ℝ] ℝ)) (𝓝 0) := begin - rw [real.cocompact_eq, tendsto_sup], - exact ⟨tendsto_integral_mul_exp_at_bot_of_continuous_compact_support hf1 hf2, - tendsto_integral_mul_exp_at_top_of_continuous_compact_support hf1 hf2⟩ + obtain ⟨C, C_ne_zero, C_ne_top, hC⟩ := μ.is_add_haar_measure_eq_smul_is_add_haar_measure volume, + rw hC, + simp_rw integral_smul_measure, + rw ←(smul_zero _ : C.to_real • (0 : E) = 0), + apply tendsto.const_smul, + let A := (inner_product_space.to_dual ℝ V).symm, + have : (λ w : V →L[ℝ] ℝ, ∫ v, e[-w v] • f v) = (λ w : V, ∫ v, e[-⟪v, w⟫] • f v) ∘ A, + { ext1 w, + congr' 1 with v : 1, + rw [←inner_conj_symm, is_R_or_C.conj_to_real, inner_product_space.to_dual_symm_apply, + real.fourier_char_apply], }, + rw this, + exact (tendsto_integral_exp_inner_smul_cocompact f).comp + A.to_homeomorph.to_cocompact_map.cocompact_tendsto', end -open_locale fourier_transform +end inner_product_space -/-- Riemann-Lebesgue lemma for continuous compactly-supported functions: the Fourier transform -tends to 0 at infinity. -/ -lemma real.fourier_integral_zero_at_infty_of_continuous_compact_support - (hc : continuous f) (hs : has_compact_support f) : - tendsto (𝓕 f) (cocompact ℝ) (𝓝 0) := +section no_inner_product + +variables + (f) [add_comm_group V] [topological_space V] [topological_add_group V] [t2_space V] + [measurable_space V] [borel_space V] + [module ℝ V] [has_continuous_smul ℝ V] [finite_dimensional ℝ V] + [complete_space E] + +/-- Riemann-Lebesgue lemma for functions on a finite-dimensional real vector space, formulated via +dual space. -/ +theorem tendsto_integral_exp_smul_cocompact (μ : measure V) [μ.is_add_haar_measure] : + tendsto (λ w : V →L[ℝ] ℝ, ∫ v, e[-w v] • f v ∂μ) (cocompact (V →L[ℝ] ℝ)) (𝓝 0) := begin - refine ((zero_at_infty_integral_mul_exp_of_continuous_compact_support hc hs).comp - (tendsto_cocompact_mul_left₀ - (mul_ne_zero (neg_ne_zero.mpr two_ne_zero) real.pi_pos.ne'))).congr (λ w, _), - rw [real.fourier_integral_eq_integral_exp_smul, function.comp_app], - congr' 1 with x:1, - ring_nf, + -- We have already proved the result for inner-product spaces, formulated in a way which doesn't + -- refer to the inner product. So we choose an arbitrary inner-product space isomorphic to V + -- and port the result over from there. + let V' := euclidean_space ℝ (fin (finrank ℝ V)), + have A : V ≃L[ℝ] V' := to_euclidean, + borelize V', + -- various equivs derived from A + let Aₘ : measurable_equiv V V' := A.to_homeomorph.to_measurable_equiv, + -- isomorphism between duals derived from A -- need to do continuity as a separate step in order + -- to apply `linear_map.continuous_of_finite_dimensional`. + let Adualₗ : (V →L[ℝ] ℝ) ≃ₗ[ℝ] (V' →L[ℝ] ℝ) := + { to_fun := λ t, t.comp A.symm.to_continuous_linear_map, + inv_fun := λ t, t.comp A.to_continuous_linear_map, + map_add' := by + { intros t s, ext1 v, simp only [continuous_linear_map.coe_comp', function.comp_app, + continuous_linear_map.add_apply] }, + map_smul' := by + { intros x f, ext1 v, simp only [ring_hom.id_apply, continuous_linear_map.coe_comp', + function.comp_app, continuous_linear_map.smul_apply] }, + left_inv := by + { intro w, ext1 v, simp only [continuous_linear_equiv.coe_def_rev, + continuous_linear_map.coe_comp', continuous_linear_equiv.coe_coe, + function.comp_app, continuous_linear_equiv.symm_apply_apply] }, + right_inv := by + { intro w, ext1 v, simp only [continuous_linear_equiv.coe_def_rev, + continuous_linear_map.coe_comp', continuous_linear_equiv.coe_coe, + function.comp_app, continuous_linear_equiv.apply_symm_apply] }, }, + let Adual : (V →L[ℝ] ℝ) ≃L[ℝ] (V' →L[ℝ] ℝ) := + { continuous_to_fun := Adualₗ.to_linear_map.continuous_of_finite_dimensional, + continuous_inv_fun := Adualₗ.symm.to_linear_map.continuous_of_finite_dimensional, + .. Adualₗ }, + haveI : (μ.map Aₘ).is_add_haar_measure, + from measure.map_continuous_linear_equiv.is_add_haar_measure _ A, + convert (tendsto_integral_exp_smul_cocompact_of_inner_product (f ∘ A.symm) (μ.map Aₘ)).comp + Adual.to_homeomorph.to_cocompact_map.cocompact_tendsto', + ext1 w, + rw [function.comp_app, integral_map_equiv], + congr' 1 with v : 1, + congr; + exact (continuous_linear_equiv.symm_apply_apply A v).symm, end -end continuous_compact_support +/-- The Riemann-Lebesgue lemma, formulated in terms of `vector_fourier.fourier_integral` (with the +pairing in the definition of `fourier_integral` taken to be the canonical pairing between `V` and +its dual space). -/ +theorem real.zero_at_infty_vector_fourier_integral (μ : measure V) [μ.is_add_haar_measure] : + tendsto (vector_fourier.fourier_integral e μ (top_dual_pairing ℝ V).flip f) + (cocompact (V →L[ℝ] ℝ)) (𝓝 0) := +tendsto_integral_exp_smul_cocompact f μ + +end no_inner_product diff --git a/src/analysis/inner_product_space/euclidean_dist.lean b/src/analysis/inner_product_space/euclidean_dist.lean index 1f921e2243159..d8a0b8fb79283 100644 --- a/src/analysis/inner_product_space/euclidean_dist.lean +++ b/src/analysis/inner_product_space/euclidean_dist.lean @@ -11,8 +11,9 @@ import analysis.inner_product_space.pi_L2 When we define a smooth bump function on a normed space, it is useful to have a smooth distance on the space. Since the default distance is not guaranteed to be smooth, we define `to_euclidean` to be -an equivalence between a finite dimensional normed space and the standard Euclidean space of the -same dimension. Then we define `euclidean.dist x y = dist (to_euclidean x) (to_euclidean y)` and +an equivalence between a finite dimensional topological vector space and the standard Euclidean +space of the same dimension. +Then we define `euclidean.dist x y = dist (to_euclidean x) (to_euclidean y)` and provide some definitions (`euclidean.ball`, `euclidean.closed_ball`) and simple lemmas about this distance. This way we hide the usage of `to_euclidean` behind an API. -/ @@ -20,13 +21,17 @@ distance. This way we hide the usage of `to_euclidean` behind an API. open_locale topology open set -variables {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] [finite_dimensional ℝ E] +variables {E : Type*} +[add_comm_group E] [topological_space E] [topological_add_group E] [t2_space E] +[module ℝ E] [has_continuous_smul ℝ E] [finite_dimensional ℝ E] + noncomputable theory +open finite_dimensional /-- If `E` is a finite dimensional space over `ℝ`, then `to_euclidean` is a continuous `ℝ`-linear equivalence between `E` and the Euclidean space of the same dimension. -/ -def to_euclidean : E ≃L[ℝ] euclidean_space ℝ (fin $ finite_dimensional.finrank ℝ E) := +def to_euclidean : E ≃L[ℝ] euclidean_space ℝ (fin $ finrank ℝ E) := continuous_linear_equiv.of_finrank_eq finrank_euclidean_space_fin.symm namespace euclidean @@ -87,7 +92,7 @@ end lemma nhds_basis_closed_ball {x : E} : (𝓝 x).has_basis (λ r : ℝ, 0 < r) (closed_ball x) := begin - rw [to_euclidean.to_homeomorph.nhds_eq_comap], + rw [to_euclidean.to_homeomorph.nhds_eq_comap x], exact metric.nhds_basis_closed_ball.comap _ end @@ -97,7 +102,7 @@ nhds_basis_closed_ball.mem_of_mem hr lemma nhds_basis_ball {x : E} : (𝓝 x).has_basis (λ r : ℝ, 0 < r) (ball x) := begin - rw [to_euclidean.to_homeomorph.nhds_eq_comap], + rw [to_euclidean.to_homeomorph.nhds_eq_comap x], exact metric.nhds_basis_ball.comap _ end @@ -106,7 +111,9 @@ nhds_basis_ball.mem_of_mem hr end euclidean -variables {F : Type*} [normed_add_comm_group F] [normed_space ℝ F] {f g : F → E} {n : ℕ∞} +variables {F : Type*} [normed_add_comm_group F] [normed_space ℝ F] + {G : Type*} [normed_add_comm_group G] [normed_space ℝ G] [finite_dimensional ℝ G] + {f g : F → G} {n : ℕ∞} lemma cont_diff.euclidean_dist (hf : cont_diff ℝ n f) (hg : cont_diff ℝ n g) (h : ∀ x, f x ≠ g x) : @@ -114,6 +121,6 @@ lemma cont_diff.euclidean_dist (hf : cont_diff ℝ n f) (hg : cont_diff ℝ n g) begin simp only [euclidean.dist], apply @cont_diff.dist ℝ, - exacts [(@to_euclidean E _ _ _).cont_diff.comp hf, - (@to_euclidean E _ _ _).cont_diff.comp hg, λ x, to_euclidean.injective.ne (h x)] + exacts [(@to_euclidean G _ _ _ _ _ _ _).cont_diff.comp hf, + (@to_euclidean G _ _ _ _ _ _ _).cont_diff.comp hg, λ x, to_euclidean.injective.ne (h x)] end diff --git a/src/analysis/normed_space/finite_dimension.lean b/src/analysis/normed_space/finite_dimension.lean index 29b0b0340edf1..b24cf93f7ea0b 100644 --- a/src/analysis/normed_space/finite_dimension.lean +++ b/src/analysis/normed_space/finite_dimension.lean @@ -252,60 +252,6 @@ begin exact is_open_set_of_linear_independent.preimage this end -/-- Two finite-dimensional normed spaces are continuously linearly equivalent if they have the same -(finite) dimension. -/ -theorem finite_dimensional.nonempty_continuous_linear_equiv_of_finrank_eq - [finite_dimensional 𝕜 E] [finite_dimensional 𝕜 F] (cond : finrank 𝕜 E = finrank 𝕜 F) : - nonempty (E ≃L[𝕜] F) := -(nonempty_linear_equiv_of_finrank_eq cond).map linear_equiv.to_continuous_linear_equiv - -/-- Two finite-dimensional normed spaces are continuously linearly equivalent if and only if they -have the same (finite) dimension. -/ -theorem finite_dimensional.nonempty_continuous_linear_equiv_iff_finrank_eq - [finite_dimensional 𝕜 E] [finite_dimensional 𝕜 F] : - nonempty (E ≃L[𝕜] F) ↔ finrank 𝕜 E = finrank 𝕜 F := -⟨ λ ⟨h⟩, h.to_linear_equiv.finrank_eq, - λ h, finite_dimensional.nonempty_continuous_linear_equiv_of_finrank_eq h ⟩ - -/-- A continuous linear equivalence between two finite-dimensional normed spaces of the same -(finite) dimension. -/ -def continuous_linear_equiv.of_finrank_eq [finite_dimensional 𝕜 E] [finite_dimensional 𝕜 F] - (cond : finrank 𝕜 E = finrank 𝕜 F) : - E ≃L[𝕜] F := -(linear_equiv.of_finrank_eq E F cond).to_continuous_linear_equiv - -variables {ι : Type*} [fintype ι] - -/-- Construct a continuous linear map given the value at a finite basis. -/ -def basis.constrL (v : basis ι 𝕜 E) (f : ι → F) : - E →L[𝕜] F := -by haveI : finite_dimensional 𝕜 E := finite_dimensional.of_fintype_basis v; - exact (v.constr 𝕜 f).to_continuous_linear_map - -@[simp, norm_cast] lemma basis.coe_constrL (v : basis ι 𝕜 E) (f : ι → F) : - (v.constrL f : E →ₗ[𝕜] F) = v.constr 𝕜 f := rfl - -/-- The continuous linear equivalence between a vector space over `𝕜` with a finite basis and -functions from its basis indexing type to `𝕜`. -/ -def basis.equiv_funL (v : basis ι 𝕜 E) : E ≃L[𝕜] (ι → 𝕜) := -{ continuous_to_fun := begin - haveI : finite_dimensional 𝕜 E := finite_dimensional.of_fintype_basis v, - exact v.equiv_fun.to_linear_map.continuous_of_finite_dimensional, - end, - continuous_inv_fun := begin - change continuous v.equiv_fun.symm.to_fun, - exact v.equiv_fun.symm.to_linear_map.continuous_of_finite_dimensional, - 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 𝕜 _ _ - -@[simp] lemma basis.constrL_basis (v : basis ι 𝕜 E) (f : ι → F) (i : ι) : - (v.constrL f) (v i) = f i := -v.constr_basis 𝕜 _ _ - lemma basis.op_nnnorm_le {ι : Type*} [fintype ι] (v : basis ι 𝕜 E) {u : E →L[𝕜] F} (M : ℝ≥0) (hu : ∀ i, ‖u (v i)‖₊ ≤ M) : ‖u‖₊ ≤ fintype.card ι • ‖v.equiv_funL.to_continuous_linear_map‖₊ * M := diff --git a/src/topology/algebra/module/finite_dimension.lean b/src/topology/algebra/module/finite_dimension.lean index ad3e32a2cbd95..e5af4572eaa6c 100644 --- a/src/topology/algebra/module/finite_dimension.lean +++ b/src/topology/algebra/module/finite_dimension.lean @@ -331,10 +331,12 @@ instance can_lift_continuous_linear_map : can_lift (E →ₗ[𝕜] F) (E →L[ end linear_map -namespace linear_equiv +section variables [t2_space E] [t2_space F] [finite_dimensional 𝕜 E] +namespace linear_equiv + /-- The continuous linear equivalence induced by a linear equivalence on a finite dimensional space. -/ def to_continuous_linear_equiv (e : E ≃ₗ[𝕜] F) : E ≃L[𝕜] F := @@ -371,6 +373,65 @@ instance can_lift_continuous_linear_equiv : end linear_equiv +variable [finite_dimensional 𝕜 F] + +/-- Two finite-dimensional topological vector spaces over a complete normed field are continuously +linearly equivalent if they have the same (finite) dimension. -/ +theorem finite_dimensional.nonempty_continuous_linear_equiv_of_finrank_eq + (cond : finrank 𝕜 E = finrank 𝕜 F) : nonempty (E ≃L[𝕜] F) := +(nonempty_linear_equiv_of_finrank_eq cond).map linear_equiv.to_continuous_linear_equiv + +/-- Two finite-dimensional topological vector spaces over a complete normed field are continuously +linearly equivalent if and only if they have the same (finite) dimension. -/ +theorem finite_dimensional.nonempty_continuous_linear_equiv_iff_finrank_eq : + nonempty (E ≃L[𝕜] F) ↔ finrank 𝕜 E = finrank 𝕜 F := +⟨ λ ⟨h⟩, h.to_linear_equiv.finrank_eq, + λ h, finite_dimensional.nonempty_continuous_linear_equiv_of_finrank_eq h ⟩ + +/-- A continuous linear equivalence between two finite-dimensional topological vector spaces over a +complete normed field of the same (finite) dimension. -/ +def continuous_linear_equiv.of_finrank_eq + (cond : finrank 𝕜 E = finrank 𝕜 F) : E ≃L[𝕜] F := +(linear_equiv.of_finrank_eq E F cond).to_continuous_linear_equiv + +end + +namespace basis + +variables {ι : Type*} [fintype ι] [t2_space E] + +/-- Construct a continuous linear map given the value at a finite basis. -/ +def constrL (v : basis ι 𝕜 E) (f : ι → F) : + E →L[𝕜] F := +by haveI : finite_dimensional 𝕜 E := finite_dimensional.of_fintype_basis v; + exact (v.constr 𝕜 f).to_continuous_linear_map + +@[simp, norm_cast] lemma coe_constrL (v : basis ι 𝕜 E) (f : ι → F) : + (v.constrL f : E →ₗ[𝕜] F) = v.constr 𝕜 f := rfl + +/-- The continuous linear equivalence between a vector space over `𝕜` with a finite basis and +functions from its basis indexing type to `𝕜`. -/ +def equiv_funL (v : basis ι 𝕜 E) : E ≃L[𝕜] (ι → 𝕜) := +{ continuous_to_fun := begin + haveI : finite_dimensional 𝕜 E := finite_dimensional.of_fintype_basis v, + exact v.equiv_fun.to_linear_map.continuous_of_finite_dimensional, + end, + continuous_inv_fun := begin + change continuous v.equiv_fun.symm.to_fun, + exact v.equiv_fun.symm.to_linear_map.continuous_of_finite_dimensional, + end, + ..v.equiv_fun } + +@[simp] lemma 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 𝕜 _ _ + +@[simp] lemma constrL_basis (v : basis ι 𝕜 E) (f : ι → F) (i : ι) : + (v.constrL f) (v i) = f i := +v.constr_basis 𝕜 _ _ + +end basis + namespace continuous_linear_map variables [t2_space E] [finite_dimensional 𝕜 E] From e42cfdb03b7902f8787a1eb552cb8f77766b45b9 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sun, 30 Apr 2023 20:09:10 +0000 Subject: [PATCH 0914/1291] chore(linear_algebra/matrix/to_linear_equiv): golf using existing morphisms (#18895) Also changes the statement of `matrix.to_linear_equiv'_symm_apply` to be `rfl`, since it's easy to change to another spelling at the call site. --- src/linear_algebra/matrix/to_linear_equiv.lean | 15 +++++---------- 1 file changed, 5 insertions(+), 10 deletions(-) diff --git a/src/linear_algebra/matrix/to_linear_equiv.lean b/src/linear_algebra/matrix/to_linear_equiv.lean index f4a54f3bd0f7f..18fe1120a77b2 100644 --- a/src/linear_algebra/matrix/to_linear_equiv.lean +++ b/src/linear_algebra/matrix/to_linear_equiv.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Patrick Massot, Casper Putz, Anne Baanen -/ import linear_algebra.finite_dimensional +import linear_algebra.matrix.general_linear_group import linear_algebra.matrix.nondegenerate import linear_algebra.matrix.nonsingular_inverse import linear_algebra.matrix.to_lin @@ -46,21 +47,15 @@ variables [decidable_eq n] See `matrix.to_linear_equiv` for the same map on arbitrary modules. -/ def to_linear_equiv' (P : matrix n n R) (h : invertible P) : (n → R) ≃ₗ[R] (n → R) := -{ inv_fun := (⅟P).to_lin', - left_inv := λ v, - show ((⅟P).to_lin'.comp P.to_lin') v = v, - by rw [← matrix.to_lin'_mul, P.inv_of_mul_self, matrix.to_lin'_one, linear_map.id_apply], - right_inv := λ v, - show (P.to_lin'.comp (⅟P).to_lin') v = v, - by rw [← matrix.to_lin'_mul, P.mul_inv_of_self, matrix.to_lin'_one, linear_map.id_apply], - ..P.to_lin' } +general_linear_group.general_linear_equiv _ _ $ + matrix.general_linear_group.to_linear $ unit_of_invertible P @[simp] lemma to_linear_equiv'_apply (P : matrix n n R) (h : invertible P) : (↑(P.to_linear_equiv' h) : module.End R (n → R)) = P.to_lin' := rfl @[simp] lemma to_linear_equiv'_symm_apply (P : matrix n n R) (h : invertible P) : - (↑(P.to_linear_equiv' h).symm : module.End R (n → R)) = P⁻¹.to_lin' := -show (⅟P).to_lin' = _, from congr_arg _ P.inv_of_eq_nonsing_inv + (↑(P.to_linear_equiv' h).symm : module.End R (n → R)) = (⅟P).to_lin' := +rfl end to_linear_equiv' From 8b8ba04e2f326f3f7cf24ad129beda58531ada61 Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Sun, 30 Apr 2023 20:09:11 +0000 Subject: [PATCH 0915/1291] feat (measure_theory/integral): lemmas `integral_Ioi_eq_integral_Ici` etc (#18899) We already have some lemmas relating integrals over open vs. closed intervals, but some were missing (notably half-infinite integrals). This should now be a full set. Co-authored-by: sgouezel --- .../integral/integrable_on.lean | 97 +++++++++++++++++++ .../integral/interval_integral.lean | 45 --------- src/measure_theory/integral/set_integral.lean | 59 +++++++++-- 3 files changed, 150 insertions(+), 51 deletions(-) diff --git a/src/measure_theory/integral/integrable_on.lean b/src/measure_theory/integral/integrable_on.lean index fb5f30403ec4c..41d7a83a9b1e8 100644 --- a/src/measure_theory/integral/integrable_on.lean +++ b/src/measure_theory/integral/integrable_on.lean @@ -563,3 +563,100 @@ lemma continuous_on.strongly_measurable_at_filter_nhds_within {α β : Type*} [m (hf : continuous_on f s) (hs : measurable_set s) (x : α) : strongly_measurable_at_filter f (𝓝[s] x) μ := ⟨s, self_mem_nhds_within, hf.ae_strongly_measurable hs⟩ + +/-! ### Lemmas about adding and removing interval boundaries + +The primed lemmas take explicit arguments about the measure being finite at the endpoint, while +the unprimed ones use `[has_no_atoms μ]`. +-/ + +section partial_order + +variables [partial_order α] [measurable_singleton_class α] + {f : α → E} {μ : measure α} {a b : α} + +lemma integrable_on_Icc_iff_integrable_on_Ioc' (ha : μ {a} ≠ ∞) : + integrable_on f (Icc a b) μ ↔ integrable_on f (Ioc a b) μ := +begin + by_cases hab : a ≤ b, + { rw [←Ioc_union_left hab, integrable_on_union, eq_true_intro + (integrable_on_singleton_iff.mpr $ or.inr ha.lt_top), and_true] }, + { rw [Icc_eq_empty hab, Ioc_eq_empty], + contrapose! hab, + exact hab.le } +end + +lemma integrable_on_Icc_iff_integrable_on_Ico' (hb : μ {b} ≠ ∞) : + integrable_on f (Icc a b) μ ↔ integrable_on f (Ico a b) μ := +begin + by_cases hab : a ≤ b, + { rw [←Ico_union_right hab, integrable_on_union, eq_true_intro + (integrable_on_singleton_iff.mpr $ or.inr hb.lt_top), and_true] }, + { rw [Icc_eq_empty hab, Ico_eq_empty], + contrapose! hab, + exact hab.le } +end + +lemma integrable_on_Ico_iff_integrable_on_Ioo' (ha : μ {a} ≠ ∞) : + integrable_on f (Ico a b) μ ↔ integrable_on f (Ioo a b) μ := +begin + by_cases hab : a < b, + { rw [←Ioo_union_left hab, integrable_on_union, eq_true_intro + (integrable_on_singleton_iff.mpr $ or.inr ha.lt_top), and_true] }, + { rw [Ioo_eq_empty hab, Ico_eq_empty hab] } +end + +lemma integrable_on_Ioc_iff_integrable_on_Ioo' (hb : μ {b} ≠ ∞) : + integrable_on f (Ioc a b) μ ↔ integrable_on f (Ioo a b) μ := +begin + by_cases hab : a < b, + { rw [←Ioo_union_right hab, integrable_on_union, eq_true_intro + (integrable_on_singleton_iff.mpr $ or.inr hb.lt_top), and_true] }, + { rw [Ioo_eq_empty hab, Ioc_eq_empty hab] } +end + +lemma integrable_on_Icc_iff_integrable_on_Ioo' (ha : μ {a} ≠ ∞) (hb : μ {b} ≠ ∞) : + integrable_on f (Icc a b) μ ↔ integrable_on f (Ioo a b) μ := +by rw [integrable_on_Icc_iff_integrable_on_Ioc' ha, integrable_on_Ioc_iff_integrable_on_Ioo' hb] + +lemma integrable_on_Ici_iff_integrable_on_Ioi' (hb : μ {b} ≠ ∞) : + integrable_on f (Ici b) μ ↔ integrable_on f (Ioi b) μ := +by rw [←Ioi_union_left, integrable_on_union, + eq_true_intro (integrable_on_singleton_iff.mpr $ or.inr hb.lt_top), and_true] + +lemma integrable_on_Iic_iff_integrable_on_Iio' (hb : μ {b} ≠ ∞) : + integrable_on f (Iic b) μ ↔ integrable_on f (Iio b) μ := +by rw [←Iio_union_right, integrable_on_union, + eq_true_intro (integrable_on_singleton_iff.mpr $ or.inr hb.lt_top), and_true] + +variables [has_no_atoms μ] + +lemma integrable_on_Icc_iff_integrable_on_Ioc : + integrable_on f (Icc a b) μ ↔ integrable_on f (Ioc a b) μ := +integrable_on_Icc_iff_integrable_on_Ioc' (by { rw measure_singleton, exact ennreal.zero_ne_top }) + +lemma integrable_on_Icc_iff_integrable_on_Ico : + integrable_on f (Icc a b) μ ↔ integrable_on f (Ico a b) μ := +integrable_on_Icc_iff_integrable_on_Ico' (by { rw measure_singleton, exact ennreal.zero_ne_top }) + +lemma integrable_on_Ico_iff_integrable_on_Ioo : + integrable_on f (Ico a b) μ ↔ integrable_on f (Ioo a b) μ := +integrable_on_Ico_iff_integrable_on_Ioo' (by { rw measure_singleton, exact ennreal.zero_ne_top }) + +lemma integrable_on_Ioc_iff_integrable_on_Ioo : + integrable_on f (Ioc a b) μ ↔ integrable_on f (Ioo a b) μ := +integrable_on_Ioc_iff_integrable_on_Ioo' (by { rw measure_singleton, exact ennreal.zero_ne_top }) + +lemma integrable_on_Icc_iff_integrable_on_Ioo : + integrable_on f (Icc a b) μ ↔ integrable_on f (Ioo a b) μ := +by rw [integrable_on_Icc_iff_integrable_on_Ioc, integrable_on_Ioc_iff_integrable_on_Ioo] + +lemma integrable_on_Ici_iff_integrable_on_Ioi : + integrable_on f (Ici b) μ ↔ integrable_on f (Ioi b) μ := +integrable_on_Ici_iff_integrable_on_Ioi' (by { rw measure_singleton, exact ennreal.zero_ne_top }) + +lemma integrable_on_Iic_iff_integrable_on_Iio : + integrable_on f (Iic b) μ ↔ integrable_on f (Iio b) μ := +integrable_on_Iic_iff_integrable_on_Iio' (by { rw measure_singleton, exact ennreal.zero_ne_top }) + +end partial_order diff --git a/src/measure_theory/integral/interval_integral.lean b/src/measure_theory/integral/interval_integral.lean index e6fb9c5c0a3db..4b7abf014c516 100644 --- a/src/measure_theory/integral/interval_integral.lean +++ b/src/measure_theory/integral/interval_integral.lean @@ -83,39 +83,6 @@ lemma interval_integrable_iff_integrable_Ioc_of_le (hab : a ≤ b) : interval_integrable f μ a b ↔ integrable_on f (Ioc a b) μ := by rw [interval_integrable_iff, uIoc_of_le hab] -lemma integrable_on_Icc_iff_integrable_on_Ioc' {f : ℝ → E} (ha : μ {a} ≠ ∞) : - integrable_on f (Icc a b) μ ↔ integrable_on f (Ioc a b) μ := -begin - cases le_or_lt a b with hab hab, - { have : Icc a b = Icc a a ∪ Ioc a b := (Icc_union_Ioc_eq_Icc le_rfl hab).symm, - rw [this, integrable_on_union], - simp [ha.lt_top] }, - { simp [hab, hab.le] }, -end - -lemma integrable_on_Icc_iff_integrable_on_Ioc [has_no_atoms μ] {f : ℝ → E} {a b : ℝ} : - integrable_on f (Icc a b) μ ↔ integrable_on f (Ioc a b) μ := -integrable_on_Icc_iff_integrable_on_Ioc' (by simp) - -lemma integrable_on_Ioc_iff_integrable_on_Ioo' - {f : ℝ → E} {a b : ℝ} (hb : μ {b} ≠ ∞) : - integrable_on f (Ioc a b) μ ↔ integrable_on f (Ioo a b) μ := -begin - cases lt_or_le a b with hab hab, - { have : Ioc a b = Ioo a b ∪ Icc b b := (Ioo_union_Icc_eq_Ioc hab le_rfl).symm, - rw [this, integrable_on_union], - simp [hb.lt_top] }, - { simp [hab] }, -end - -lemma integrable_on_Ioc_iff_integrable_on_Ioo [has_no_atoms μ] {f : ℝ → E} {a b : ℝ} : - integrable_on f (Ioc a b) μ ↔ integrable_on f (Ioo a b) μ := -integrable_on_Ioc_iff_integrable_on_Ioo' (by simp) - -lemma integrable_on_Icc_iff_integrable_on_Ioo [has_no_atoms μ] {f : ℝ → E} {a b : ℝ} : - integrable_on f (Icc a b) μ ↔ integrable_on f (Ioo a b) μ := -by rw [integrable_on_Icc_iff_integrable_on_Ioc, integrable_on_Ioc_iff_integrable_on_Ioo] - lemma interval_integrable_iff' [has_no_atoms μ] : interval_integrable f μ a b ↔ integrable_on f (uIcc a b) μ := by rw [interval_integrable_iff, ←Icc_min_max, uIoc, integrable_on_Icc_iff_integrable_on_Ioc] @@ -125,18 +92,6 @@ lemma interval_integrable_iff_integrable_Icc_of_le interval_integrable f μ a b ↔ integrable_on f (Icc a b) μ := by rw [interval_integrable_iff_integrable_Ioc_of_le hab, integrable_on_Icc_iff_integrable_on_Ioc] -lemma integrable_on_Ici_iff_integrable_on_Ioi' {f : ℝ → E} (ha : μ {a} ≠ ∞) : - integrable_on f (Ici a) μ ↔ integrable_on f (Ioi a) μ := -begin - have : Ici a = Icc a a ∪ Ioi a := (Icc_union_Ioi_eq_Ici le_rfl).symm, - rw [this, integrable_on_union], - simp [ha.lt_top] -end - -lemma integrable_on_Ici_iff_integrable_on_Ioi [has_no_atoms μ] {f : ℝ → E} : - integrable_on f (Ici a) μ ↔ integrable_on f (Ioi a) μ := -integrable_on_Ici_iff_integrable_on_Ioi' (by simp) - /-- If a function is integrable with respect to a given measure `μ` then it is interval integrable with respect to `μ` on `uIcc a b`. -/ lemma measure_theory.integrable.interval_integrable (hf : integrable f μ) : diff --git a/src/measure_theory/integral/set_integral.lean b/src/measure_theory/integral/set_integral.lean index a5d4d950ca8aa..fc4be754323a5 100644 --- a/src/measure_theory/integral/set_integral.lean +++ b/src/measure_theory/integral/set_integral.lean @@ -593,22 +593,69 @@ lemma set_integral_trim {α} {m m0 : measurable_space α} {μ : measure α} (hm ∫ x in s, f x ∂μ = ∫ x in s, f x ∂(μ.trim hm) := by rwa [integral_trim hm hf_meas, restrict_trim hm μ] -lemma integral_Icc_eq_integral_Ioc' [partial_order α] {f : α → E} {a b : α} (ha : μ {a} = 0) : +/-! ### Lemmas about adding and removing interval boundaries + +The primed lemmas take explicit arguments about the endpoint having zero measure, while the +unprimed ones use `[has_no_atoms μ]`. +-/ + +section partial_order + +variables [partial_order α] {a b : α} + +lemma integral_Icc_eq_integral_Ioc' (ha : μ {a} = 0) : ∫ t in Icc a b, f t ∂μ = ∫ t in Ioc a b, f t ∂μ := set_integral_congr_set_ae (Ioc_ae_eq_Icc' ha).symm -lemma integral_Ioc_eq_integral_Ioo' [partial_order α] {f : α → E} {a b : α} (hb : μ {b} = 0) : +lemma integral_Icc_eq_integral_Ico' (hb : μ {b} = 0) : + ∫ t in Icc a b, f t ∂μ = ∫ t in Ico a b, f t ∂μ := +set_integral_congr_set_ae (Ico_ae_eq_Icc' hb).symm + +lemma integral_Ioc_eq_integral_Ioo' (hb : μ {b} = 0) : ∫ t in Ioc a b, f t ∂μ = ∫ t in Ioo a b, f t ∂μ := set_integral_congr_set_ae (Ioo_ae_eq_Ioc' hb).symm -lemma integral_Icc_eq_integral_Ioc [partial_order α] {f : α → E} {a b : α} [has_no_atoms μ] : - ∫ t in Icc a b, f t ∂μ = ∫ t in Ioc a b, f t ∂μ := +lemma integral_Ico_eq_integral_Ioo' (ha : μ {a} = 0) : + ∫ t in Ico a b, f t ∂μ = ∫ t in Ioo a b, f t ∂μ := +set_integral_congr_set_ae (Ioo_ae_eq_Ico' ha).symm + +lemma integral_Icc_eq_integral_Ioo' (ha : μ {a} = 0) (hb : μ {b} = 0) : + ∫ t in Icc a b, f t ∂μ = ∫ t in Ioo a b, f t ∂μ := +set_integral_congr_set_ae (Ioo_ae_eq_Icc' ha hb).symm + +lemma integral_Iic_eq_integral_Iio' (ha : μ {a} = 0) : + ∫ t in Iic a, f t ∂μ = ∫ t in Iio a, f t ∂μ := +set_integral_congr_set_ae (Iio_ae_eq_Iic' ha).symm + +lemma integral_Ici_eq_integral_Ioi' (ha : μ {a} = 0) : + ∫ t in Ici a, f t ∂μ = ∫ t in Ioi a, f t ∂μ := +set_integral_congr_set_ae (Ioi_ae_eq_Ici' ha).symm + +variable [has_no_atoms μ] + +lemma integral_Icc_eq_integral_Ioc : ∫ t in Icc a b, f t ∂μ = ∫ t in Ioc a b, f t ∂μ := integral_Icc_eq_integral_Ioc' $ measure_singleton a -lemma integral_Ioc_eq_integral_Ioo [partial_order α] {f : α → E} {a b : α} [has_no_atoms μ] : - ∫ t in Ioc a b, f t ∂μ = ∫ t in Ioo a b, f t ∂μ := +lemma integral_Icc_eq_integral_Ico : ∫ t in Icc a b, f t ∂μ = ∫ t in Ico a b, f t ∂μ := +integral_Icc_eq_integral_Ico' $ measure_singleton b + +lemma integral_Ioc_eq_integral_Ioo : ∫ t in Ioc a b, f t ∂μ = ∫ t in Ioo a b, f t ∂μ := integral_Ioc_eq_integral_Ioo' $ measure_singleton b +lemma integral_Ico_eq_integral_Ioo : ∫ t in Ico a b, f t ∂μ = ∫ t in Ioo a b, f t ∂μ := +integral_Ico_eq_integral_Ioo' $ measure_singleton a + +lemma integral_Icc_eq_integral_Ioo : ∫ t in Icc a b, f t ∂μ = ∫ t in Ico a b, f t ∂μ := +by rw [integral_Icc_eq_integral_Ico, integral_Ico_eq_integral_Ioo] + +lemma integral_Iic_eq_integral_Iio : ∫ t in Iic a, f t ∂μ = ∫ t in Iio a, f t ∂μ := +integral_Iic_eq_integral_Iio' $ measure_singleton a + +lemma integral_Ici_eq_integral_Ioi : ∫ t in Ici a, f t ∂μ = ∫ t in Ioi a, f t ∂μ := +integral_Ici_eq_integral_Ioi' $ measure_singleton a + +end partial_order + end normed_add_comm_group section mono From d20a8cd58d30f2463dbf59f0c398a747854243ed Mon Sep 17 00:00:00 2001 From: Alex J Best Date: Sun, 30 Apr 2023 23:23:16 +0000 Subject: [PATCH 0916/1291] chore(.github/workflows): update gh-get-current-pr (#18763) The test that this upgrade works ok is that the awaiting-CI label is removed --- .github/workflows/bors.yml | 2 +- .github/workflows/build.yml | 2 +- .github/workflows/build.yml.in | 2 +- .github/workflows/build_fork.yml | 2 +- 4 files changed, 4 insertions(+), 4 deletions(-) diff --git a/.github/workflows/bors.yml b/.github/workflows/bors.yml index a742b9ae0d2da..15b70725de7a1 100644 --- a/.github/workflows/bors.yml +++ b/.github/workflows/bors.yml @@ -260,7 +260,7 @@ jobs: - uses: actions/checkout@v3 - id: PR - uses: 8BitJonny/gh-get-current-pr@1.2.0 + uses: 8BitJonny/gh-get-current-pr@2.2.0 # TODO: this may not work properly if the same commit is pushed to multiple branches: # https://github.com/8BitJonny/gh-get-current-pr/issues/8 with: diff --git a/.github/workflows/build.yml b/.github/workflows/build.yml index 0b696f4e052a1..34f6945efa691 100644 --- a/.github/workflows/build.yml +++ b/.github/workflows/build.yml @@ -268,7 +268,7 @@ jobs: - uses: actions/checkout@v3 - id: PR - uses: 8BitJonny/gh-get-current-pr@1.2.0 + uses: 8BitJonny/gh-get-current-pr@2.2.0 # TODO: this may not work properly if the same commit is pushed to multiple branches: # https://github.com/8BitJonny/gh-get-current-pr/issues/8 with: diff --git a/.github/workflows/build.yml.in b/.github/workflows/build.yml.in index dab918a1a7e7b..44b295b967665 100644 --- a/.github/workflows/build.yml.in +++ b/.github/workflows/build.yml.in @@ -246,7 +246,7 @@ jobs: - uses: actions/checkout@v3 - id: PR - uses: 8BitJonny/gh-get-current-pr@1.2.0 + uses: 8BitJonny/gh-get-current-pr@2.2.0 # TODO: this may not work properly if the same commit is pushed to multiple branches: # https://github.com/8BitJonny/gh-get-current-pr/issues/8 with: diff --git a/.github/workflows/build_fork.yml b/.github/workflows/build_fork.yml index 2541ccc498838..ecc2a12a2f0fd 100644 --- a/.github/workflows/build_fork.yml +++ b/.github/workflows/build_fork.yml @@ -266,7 +266,7 @@ jobs: - uses: actions/checkout@v3 - id: PR - uses: 8BitJonny/gh-get-current-pr@1.2.0 + uses: 8BitJonny/gh-get-current-pr@2.2.0 # TODO: this may not work properly if the same commit is pushed to multiple branches: # https://github.com/8BitJonny/gh-get-current-pr/issues/8 with: From 7317149f12f55affbc900fc873d0d422485122b9 Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Mon, 1 May 2023 03:53:08 +0000 Subject: [PATCH 0917/1291] chore(measure_theory/integral): split up measure_theory.integral.lebesgue (#18904) This PR shortens the 3056-line-long file `measure_theory.integral.lebesgue`, by removing a 1000-line initial segment into a new file `measure_theory.function.simple_func`. --- src/measure_theory/function/simple_func.lean | 1012 +++++++++++++++++ .../function/simple_func_dense.lean | 2 +- .../function/strongly_measurable/basic.lean | 1 + src/measure_theory/integral/lebesgue.lean | 1005 +--------------- 4 files changed, 1023 insertions(+), 997 deletions(-) create mode 100644 src/measure_theory/function/simple_func.lean diff --git a/src/measure_theory/function/simple_func.lean b/src/measure_theory/function/simple_func.lean new file mode 100644 index 0000000000000..9051f119f2976 --- /dev/null +++ b/src/measure_theory/function/simple_func.lean @@ -0,0 +1,1012 @@ +/- +Copyright (c) 2018 Mario Carneiro. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Mario Carneiro, Johannes Hölzl +-/ +import measure_theory.constructions.borel_space +import algebra.indicator_function +import algebra.support + +/-! +# Simple functions + +A function `f` from a measurable space to any type is called *simple*, if every preimage `f ⁻¹' {x}` +is measurable, and the range is finite. In this file, we define simple functions and establish their +basic properties; and we construct a sequence of simple functions approximating an arbitrary Borel +measurable function `f : α → ℝ≥0∞`. + +The theorem `measurable.ennreal_induction` shows that in order to prove something for an arbitrary +measurable function into `ℝ≥0∞`, it is sufficient to show that the property holds for (multiples of) +characteristic functions and is closed under addition and supremum of increasing sequences of +functions. +-/ +noncomputable theory +open set (hiding restrict restrict_apply) filter ennreal function (support) +open_locale classical topology big_operators nnreal ennreal measure_theory + +namespace measure_theory + +variables {α β γ δ : Type*} + +/-- A function `f` from a measurable space to any type is called *simple*, +if every preimage `f ⁻¹' {x}` is measurable, and the range is finite. This structure bundles +a function with these properties. -/ +structure {u v} simple_func (α : Type u) [measurable_space α] (β : Type v) := +(to_fun : α → β) +(measurable_set_fiber' : ∀ x, measurable_set (to_fun ⁻¹' {x})) +(finite_range' : (set.range to_fun).finite) + +local infixr ` →ₛ `:25 := simple_func + +namespace simple_func + +section measurable +variables [measurable_space α] +instance has_coe_to_fun : has_coe_to_fun (α →ₛ β) (λ _, α → β) := ⟨to_fun⟩ + +lemma coe_injective ⦃f g : α →ₛ β⦄ (H : (f : α → β) = g) : f = g := +by cases f; cases g; congr; exact H + +@[ext] theorem ext {f g : α →ₛ β} (H : ∀ a, f a = g a) : f = g := +coe_injective $ funext H + +lemma finite_range (f : α →ₛ β) : (set.range f).finite := f.finite_range' + +lemma measurable_set_fiber (f : α →ₛ β) (x : β) : measurable_set (f ⁻¹' {x}) := +f.measurable_set_fiber' x + +@[simp] lemma apply_mk (f : α → β) (h h') (x : α) : simple_func.mk f h h' x = f x := rfl + +/-- Simple function defined on the empty type. -/ +def of_is_empty [is_empty α] : α →ₛ β := +{ to_fun := is_empty_elim, + measurable_set_fiber' := λ x, subsingleton.measurable_set, + finite_range' := by simp [range_eq_empty] } + +/-- Range of a simple function `α →ₛ β` as a `finset β`. -/ +protected def range (f : α →ₛ β) : finset β := f.finite_range.to_finset + +@[simp] theorem mem_range {f : α →ₛ β} {b} : b ∈ f.range ↔ b ∈ range f := +finite.mem_to_finset _ + +theorem mem_range_self (f : α →ₛ β) (x : α) : f x ∈ f.range := mem_range.2 ⟨x, rfl⟩ + +@[simp] lemma coe_range (f : α →ₛ β) : (↑f.range : set β) = set.range f := +f.finite_range.coe_to_finset + +theorem mem_range_of_measure_ne_zero {f : α →ₛ β} {x : β} {μ : measure α} (H : μ (f ⁻¹' {x}) ≠ 0) : + x ∈ f.range := +let ⟨a, ha⟩ := nonempty_of_measure_ne_zero H in +mem_range.2 ⟨a, ha⟩ + +lemma forall_range_iff {f : α →ₛ β} {p : β → Prop} : + (∀ y ∈ f.range, p y) ↔ ∀ x, p (f x) := +by simp only [mem_range, set.forall_range_iff] + +lemma exists_range_iff {f : α →ₛ β} {p : β → Prop} : + (∃ y ∈ f.range, p y) ↔ ∃ x, p (f x) := +by simpa only [mem_range, exists_prop] using set.exists_range_iff + +lemma preimage_eq_empty_iff (f : α →ₛ β) (b : β) : f ⁻¹' {b} = ∅ ↔ b ∉ f.range := +preimage_singleton_eq_empty.trans $ not_congr mem_range.symm + +lemma exists_forall_le [nonempty β] [preorder β] [is_directed β (≤)] (f : α →ₛ β) : + ∃ C, ∀ x, f x ≤ C := +f.range.exists_le.imp $ λ C, forall_range_iff.1 + +/-- Constant function as a `simple_func`. -/ +def const (α) {β} [measurable_space α] (b : β) : α →ₛ β := +⟨λ a, b, λ x, measurable_set.const _, finite_range_const⟩ + +instance [inhabited β] : inhabited (α →ₛ β) := ⟨const _ default⟩ + +theorem const_apply (a : α) (b : β) : (const α b) a = b := rfl + +@[simp] theorem coe_const (b : β) : ⇑(const α b) = function.const α b := rfl + +@[simp] lemma range_const (α) [measurable_space α] [nonempty α] (b : β) : + (const α b).range = {b} := +finset.coe_injective $ by simp + +lemma range_const_subset (α) [measurable_space α] (b : β) : + (const α b).range ⊆ {b} := +finset.coe_subset.1 $ by simp + +lemma simple_func_bot {α} (f : @simple_func α ⊥ β) [nonempty β] : ∃ c, ∀ x, f x = c := +begin + have hf_meas := @simple_func.measurable_set_fiber α _ ⊥ f, + simp_rw measurable_space.measurable_set_bot_iff at hf_meas, + casesI is_empty_or_nonempty α, + { simp only [is_empty.forall_iff, exists_const], }, + { specialize hf_meas (f h.some), + cases hf_meas, + { exfalso, + refine set.not_mem_empty h.some _, + rw [← hf_meas, set.mem_preimage], + exact set.mem_singleton _, }, + { refine ⟨f h.some, λ x, _⟩, + have : x ∈ f ⁻¹' {f h.some}, + { rw hf_meas, exact set.mem_univ x, }, + rwa [set.mem_preimage, set.mem_singleton_iff] at this, }, }, +end + +lemma simple_func_bot' {α} [nonempty β] (f : @simple_func α ⊥ β) : + ∃ c, f = @simple_func.const α _ ⊥ c := +begin + obtain ⟨c, h_eq⟩ := simple_func_bot f, + refine ⟨c, _⟩, + ext1 x, + rw [h_eq x, simple_func.coe_const], +end + +lemma measurable_set_cut (r : α → β → Prop) (f : α →ₛ β) + (h : ∀b, measurable_set {a | r a b}) : measurable_set {a | r a (f a)} := +begin + have : {a | r a (f a)} = ⋃ b ∈ range f, {a | r a b} ∩ f ⁻¹' {b}, + { ext a, + suffices : r a (f a) ↔ ∃ i, r a (f i) ∧ f a = f i, by simpa, + exact ⟨λ h, ⟨a, ⟨h, rfl⟩⟩, λ ⟨a', ⟨h', e⟩⟩, e.symm ▸ h'⟩ }, + rw this, + exact measurable_set.bUnion f.finite_range.countable + (λ b _, measurable_set.inter (h b) (f.measurable_set_fiber _)) +end + +@[measurability] +theorem measurable_set_preimage (f : α →ₛ β) (s) : measurable_set (f ⁻¹' s) := +measurable_set_cut (λ _ b, b ∈ s) f (λ b, measurable_set.const (b ∈ s)) + +/-- A simple function is measurable -/ +@[measurability] +protected theorem measurable [measurable_space β] (f : α →ₛ β) : measurable f := +λ s _, measurable_set_preimage f s + +@[measurability] +protected theorem ae_measurable [measurable_space β] {μ : measure α} (f : α →ₛ β) : + ae_measurable f μ := +f.measurable.ae_measurable + +protected lemma sum_measure_preimage_singleton (f : α →ₛ β) {μ : measure α} (s : finset β) : + ∑ y in s, μ (f ⁻¹' {y}) = μ (f ⁻¹' ↑s) := +sum_measure_preimage_singleton _ (λ _ _, f.measurable_set_fiber _) + +lemma sum_range_measure_preimage_singleton (f : α →ₛ β) (μ : measure α) : + ∑ y in f.range, μ (f ⁻¹' {y}) = μ univ := +by rw [f.sum_measure_preimage_singleton, coe_range, preimage_range] + +/-- If-then-else as a `simple_func`. -/ +def piecewise (s : set α) (hs : measurable_set s) (f g : α →ₛ β) : α →ₛ β := +⟨s.piecewise f g, + λ x, by letI : measurable_space β := ⊤; exact + f.measurable.piecewise hs g.measurable trivial, + (f.finite_range.union g.finite_range).subset range_ite_subset⟩ + +@[simp] theorem coe_piecewise {s : set α} (hs : measurable_set s) (f g : α →ₛ β) : + ⇑(piecewise s hs f g) = s.piecewise f g := +rfl + +theorem piecewise_apply {s : set α} (hs : measurable_set s) (f g : α →ₛ β) (a) : + piecewise s hs f g a = if a ∈ s then f a else g a := +rfl + +@[simp] lemma piecewise_compl {s : set α} (hs : measurable_set sᶜ) (f g : α →ₛ β) : + piecewise sᶜ hs f g = piecewise s hs.of_compl g f := +coe_injective $ by simp [hs] + +@[simp] lemma piecewise_univ (f g : α →ₛ β) : piecewise univ measurable_set.univ f g = f := +coe_injective $ by simp + +@[simp] lemma piecewise_empty (f g : α →ₛ β) : piecewise ∅ measurable_set.empty f g = g := +coe_injective $ by simp + +lemma support_indicator [has_zero β] {s : set α} (hs : measurable_set s) (f : α →ₛ β) : + function.support (f.piecewise s hs (simple_func.const α 0)) = s ∩ function.support f := +set.support_indicator + +lemma range_indicator {s : set α} (hs : measurable_set s) + (hs_nonempty : s.nonempty) (hs_ne_univ : s ≠ univ) (x y : β) : + (piecewise s hs (const α x) (const α y)).range = {x, y} := +by simp only [← finset.coe_inj, coe_range, coe_piecewise, range_piecewise, coe_const, + finset.coe_insert, finset.coe_singleton, hs_nonempty.image_const, + (nonempty_compl.2 hs_ne_univ).image_const, singleton_union] + +lemma measurable_bind [measurable_space γ] (f : α →ₛ β) (g : β → α → γ) + (hg : ∀ b, measurable (g b)) : measurable (λ a, g (f a) a) := +λ s hs, f.measurable_set_cut (λ a b, g b a ∈ s) $ λ b, hg b hs + +/-- If `f : α →ₛ β` is a simple function and `g : β → α →ₛ γ` is a family of simple functions, +then `f.bind g` binds the first argument of `g` to `f`. In other words, `f.bind g a = g (f a) a`. -/ +def bind (f : α →ₛ β) (g : β → α →ₛ γ) : α →ₛ γ := +⟨λa, g (f a) a, + λ c, f.measurable_set_cut (λ a b, g b a = c) $ λ b, (g b).measurable_set_preimage {c}, + (f.finite_range.bUnion (λ b _, (g b).finite_range)).subset $ + by rintro _ ⟨a, rfl⟩; simp; exact ⟨a, a, rfl⟩⟩ + +@[simp] theorem bind_apply (f : α →ₛ β) (g : β → α →ₛ γ) (a) : + f.bind g a = g (f a) a := rfl + +/-- Given a function `g : β → γ` and a simple function `f : α →ₛ β`, `f.map g` return the simple + function `g ∘ f : α →ₛ γ` -/ +def map (g : β → γ) (f : α →ₛ β) : α →ₛ γ := bind f (const α ∘ g) + +theorem map_apply (g : β → γ) (f : α →ₛ β) (a) : f.map g a = g (f a) := rfl + +theorem map_map (g : β → γ) (h: γ → δ) (f : α →ₛ β) : (f.map g).map h = f.map (h ∘ g) := rfl + +@[simp] theorem coe_map (g : β → γ) (f : α →ₛ β) : (f.map g : α → γ) = g ∘ f := rfl + +@[simp] theorem range_map [decidable_eq γ] (g : β → γ) (f : α →ₛ β) : + (f.map g).range = f.range.image g := +finset.coe_injective $ by simp only [coe_range, coe_map, finset.coe_image, range_comp] + +@[simp] theorem map_const (g : β → γ) (b : β) : (const α b).map g = const α (g b) := rfl + +lemma map_preimage (f : α →ₛ β) (g : β → γ) (s : set γ) : + (f.map g) ⁻¹' s = f ⁻¹' ↑(f.range.filter (λb, g b ∈ s)) := +by { simp only [coe_range, sep_mem_eq, set.mem_range, function.comp_app, coe_map, finset.coe_filter, + ← mem_preimage, inter_comm, preimage_inter_range], apply preimage_comp } + +lemma map_preimage_singleton (f : α →ₛ β) (g : β → γ) (c : γ) : + (f.map g) ⁻¹' {c} = f ⁻¹' ↑(f.range.filter (λ b, g b = c)) := +map_preimage _ _ _ + +/-- Composition of a `simple_fun` and a measurable function is a `simple_func`. -/ +def comp [measurable_space β] (f : β →ₛ γ) (g : α → β) (hgm : measurable g) : α →ₛ γ := +{ to_fun := f ∘ g, + finite_range' := f.finite_range.subset $ set.range_comp_subset_range _ _, + measurable_set_fiber' := λ z, hgm (f.measurable_set_fiber z) } + +@[simp] lemma coe_comp [measurable_space β] (f : β →ₛ γ) {g : α → β} (hgm : measurable g) : + ⇑(f.comp g hgm) = f ∘ g := +rfl + +lemma range_comp_subset_range [measurable_space β] (f : β →ₛ γ) {g : α → β} (hgm : measurable g) : + (f.comp g hgm).range ⊆ f.range := +finset.coe_subset.1 $ by simp only [coe_range, coe_comp, set.range_comp_subset_range] + +/-- Extend a `simple_func` along a measurable embedding: `f₁.extend g hg f₂` is the function +`F : β →ₛ γ` such that `F ∘ g = f₁` and `F y = f₂ y` whenever `y ∉ range g`. -/ +def extend [measurable_space β] (f₁ : α →ₛ γ) (g : α → β) + (hg : measurable_embedding g) (f₂ : β →ₛ γ) : β →ₛ γ := +{ to_fun := function.extend g f₁ f₂, + finite_range' := (f₁.finite_range.union $ f₂.finite_range.subset + (image_subset_range _ _)).subset (range_extend_subset _ _ _), + measurable_set_fiber' := + begin + letI : measurable_space γ := ⊤, haveI : measurable_singleton_class γ := ⟨λ _, trivial⟩, + exact λ x, hg.measurable_extend f₁.measurable f₂.measurable (measurable_set_singleton _) + end } + +@[simp] lemma extend_apply [measurable_space β] (f₁ : α →ₛ γ) {g : α → β} + (hg : measurable_embedding g) (f₂ : β →ₛ γ) (x : α) : (f₁.extend g hg f₂) (g x) = f₁ x := +hg.injective.extend_apply _ _ _ + +@[simp] lemma extend_apply' [measurable_space β] (f₁ : α →ₛ γ) {g : α → β} + (hg : measurable_embedding g) (f₂ : β →ₛ γ) {y : β} (h : ¬∃ x, g x = y) : + (f₁.extend g hg f₂) y = f₂ y := +function.extend_apply' _ _ _ h + +@[simp] lemma extend_comp_eq' [measurable_space β] (f₁ : α →ₛ γ) {g : α → β} + (hg : measurable_embedding g) (f₂ : β →ₛ γ) : (f₁.extend g hg f₂) ∘ g = f₁ := +funext $ λ x, extend_apply _ _ _ _ + +@[simp] lemma extend_comp_eq [measurable_space β] (f₁ : α →ₛ γ) {g : α → β} + (hg : measurable_embedding g) (f₂ : β →ₛ γ) : (f₁.extend g hg f₂).comp g hg.measurable = f₁ := +coe_injective $ extend_comp_eq' _ _ _ + +/-- If `f` is a simple function taking values in `β → γ` and `g` is another simple function +with the same domain and codomain `β`, then `f.seq g = f a (g a)`. -/ +def seq (f : α →ₛ (β → γ)) (g : α →ₛ β) : α →ₛ γ := f.bind (λf, g.map f) + +@[simp] lemma seq_apply (f : α →ₛ (β → γ)) (g : α →ₛ β) (a : α) : f.seq g a = f a (g a) := rfl + +/-- Combine two simple functions `f : α →ₛ β` and `g : α →ₛ β` +into `λ a, (f a, g a)`. -/ +def pair (f : α →ₛ β) (g : α →ₛ γ) : α →ₛ (β × γ) := (f.map prod.mk).seq g + +@[simp] lemma pair_apply (f : α →ₛ β) (g : α →ₛ γ) (a) : pair f g a = (f a, g a) := rfl + +lemma pair_preimage (f : α →ₛ β) (g : α →ₛ γ) (s : set β) (t : set γ) : + pair f g ⁻¹' s ×ˢ t = (f ⁻¹' s) ∩ (g ⁻¹' t) := rfl + +/- A special form of `pair_preimage` -/ +lemma pair_preimage_singleton (f : α →ₛ β) (g : α →ₛ γ) (b : β) (c : γ) : + (pair f g) ⁻¹' {(b, c)} = (f ⁻¹' {b}) ∩ (g ⁻¹' {c}) := +by { rw ← singleton_prod_singleton, exact pair_preimage _ _ _ _ } + +theorem bind_const (f : α →ₛ β) : f.bind (const α) = f := by ext; simp + +@[to_additive] instance [has_one β] : has_one (α →ₛ β) := ⟨const α 1⟩ +@[to_additive] instance [has_mul β] : has_mul (α →ₛ β) := ⟨λf g, (f.map (*)).seq g⟩ +@[to_additive] instance [has_div β] : has_div (α →ₛ β) := ⟨λf g, (f.map (/)).seq g⟩ +@[to_additive] instance [has_inv β] : has_inv (α →ₛ β) := ⟨λf, f.map (has_inv.inv)⟩ +instance [has_sup β] : has_sup (α →ₛ β) := ⟨λf g, (f.map (⊔)).seq g⟩ +instance [has_inf β] : has_inf (α →ₛ β) := ⟨λf g, (f.map (⊓)).seq g⟩ +instance [has_le β] : has_le (α →ₛ β) := ⟨λf g, ∀a, f a ≤ g a⟩ + +@[simp, to_additive] lemma const_one [has_one β] : const α (1 : β) = 1 := rfl + +@[simp, norm_cast, to_additive] lemma coe_one [has_one β] : ⇑(1 : α →ₛ β) = 1 := rfl +@[simp, norm_cast, to_additive] lemma coe_mul [has_mul β] (f g : α →ₛ β) : ⇑(f * g) = f * g := rfl +@[simp, norm_cast, to_additive] lemma coe_inv [has_inv β] (f : α →ₛ β) : ⇑(f⁻¹) = f⁻¹ := rfl +@[simp, norm_cast, to_additive] lemma coe_div [has_div β] (f g : α →ₛ β) : ⇑(f / g) = f / g := rfl +@[simp, norm_cast] lemma coe_le [preorder β] {f g : α →ₛ β} : (f : α → β) ≤ g ↔ f ≤ g := iff.rfl +@[simp, norm_cast] lemma coe_sup [has_sup β] (f g : α →ₛ β) : ⇑(f ⊔ g) = f ⊔ g := rfl +@[simp, norm_cast] lemma coe_inf [has_inf β] (f g : α →ₛ β) : ⇑(f ⊓ g) = f ⊓ g := rfl + +@[to_additive] lemma mul_apply [has_mul β] (f g : α →ₛ β) (a : α) : (f * g) a = f a * g a := rfl +@[to_additive] lemma div_apply [has_div β] (f g : α →ₛ β) (x : α) : (f / g) x = f x / g x := rfl +@[to_additive] lemma inv_apply [has_inv β] (f : α →ₛ β) (x : α) : f⁻¹ x = (f x)⁻¹ := rfl +lemma sup_apply [has_sup β] (f g : α →ₛ β) (a : α) : (f ⊔ g) a = f a ⊔ g a := rfl +lemma inf_apply [has_inf β] (f g : α →ₛ β) (a : α) : (f ⊓ g) a = f a ⊓ g a := rfl + +@[simp, to_additive] lemma range_one [nonempty α] [has_one β] : (1 : α →ₛ β).range = {1} := +finset.ext $ λ x, by simp [eq_comm] + +@[simp] lemma range_eq_empty_of_is_empty {β} [hα : is_empty α] (f : α →ₛ β) : + f.range = ∅ := +begin + rw ← finset.not_nonempty_iff_eq_empty, + by_contra, + obtain ⟨y, hy_mem⟩ := h, + rw [simple_func.mem_range, set.mem_range] at hy_mem, + obtain ⟨x, hxy⟩ := hy_mem, + rw is_empty_iff at hα, + exact hα x, +end + +lemma eq_zero_of_mem_range_zero [has_zero β] : ∀ {y : β}, y ∈ (0 : α →ₛ β).range → y = 0 := +forall_range_iff.2 $ λ x, rfl + +@[to_additive] +lemma mul_eq_map₂ [has_mul β] (f g : α →ₛ β) : f * g = (pair f g).map (λp:β×β, p.1 * p.2) := rfl + +lemma sup_eq_map₂ [has_sup β] (f g : α →ₛ β) : f ⊔ g = (pair f g).map (λp:β×β, p.1 ⊔ p.2) := rfl + +@[to_additive] +lemma const_mul_eq_map [has_mul β] (f : α →ₛ β) (b : β) : const α b * f = f.map (λa, b * a) := rfl + +@[to_additive] +theorem map_mul [has_mul β] [has_mul γ] {g : β → γ} + (hg : ∀ x y, g (x * y) = g x * g y) (f₁ f₂ : α →ₛ β) : (f₁ * f₂).map g = f₁.map g * f₂.map g := +ext $ λ x, hg _ _ + +variables {K : Type*} + +instance [has_smul K β] : has_smul K (α →ₛ β) := ⟨λ k f, f.map ((•) k)⟩ +@[simp] lemma coe_smul [has_smul K β] (c : K) (f : α →ₛ β) : ⇑(c • f) = c • f := rfl +lemma smul_apply [has_smul K β] (k : K) (f : α →ₛ β) (a : α) : (k • f) a = k • f a := rfl + +instance has_nat_pow [monoid β] : has_pow (α →ₛ β) ℕ := ⟨λ f n, f.map (^ n)⟩ +@[simp] lemma coe_pow [monoid β] (f : α →ₛ β) (n : ℕ) : ⇑(f ^ n) = f ^ n := rfl +lemma pow_apply [monoid β] (n : ℕ) (f : α →ₛ β) (a : α) : (f ^ n) a = f a ^ n := rfl + +instance has_int_pow [div_inv_monoid β] : has_pow (α →ₛ β) ℤ := ⟨λ f n, f.map (^ n)⟩ +@[simp] lemma coe_zpow [div_inv_monoid β] (f : α →ₛ β) (z : ℤ) : ⇑(f ^ z) = f ^ z := rfl +lemma zpow_apply [div_inv_monoid β] (z : ℤ) (f : α →ₛ β) (a : α) : (f ^ z) a = f a ^ z := rfl + +-- TODO: work out how to generate these instances with `to_additive`, which gets confused by the +-- argument order swap between `coe_smul` and `coe_pow`. +section additive + +instance [add_monoid β] : add_monoid (α →ₛ β) := +function.injective.add_monoid (λ f, show α → β, from f) coe_injective coe_zero coe_add + (λ _ _, coe_smul _ _) + +instance [add_comm_monoid β] : add_comm_monoid (α →ₛ β) := +function.injective.add_comm_monoid (λ f, show α → β, from f) coe_injective coe_zero coe_add + (λ _ _, coe_smul _ _) + +instance [add_group β] : add_group (α →ₛ β) := +function.injective.add_group (λ f, show α → β, from f) coe_injective + coe_zero coe_add coe_neg coe_sub (λ _ _, coe_smul _ _) (λ _ _, coe_smul _ _) + +instance [add_comm_group β] : add_comm_group (α →ₛ β) := +function.injective.add_comm_group (λ f, show α → β, from f) coe_injective + coe_zero coe_add coe_neg coe_sub (λ _ _, coe_smul _ _) (λ _ _, coe_smul _ _) + +end additive + +@[to_additive] instance [monoid β] : monoid (α →ₛ β) := +function.injective.monoid (λ f, show α → β, from f) coe_injective coe_one coe_mul coe_pow + +@[to_additive] instance [comm_monoid β] : comm_monoid (α →ₛ β) := +function.injective.comm_monoid (λ f, show α → β, from f) coe_injective coe_one coe_mul coe_pow + +@[to_additive] instance [group β] : group (α →ₛ β) := +function.injective.group (λ f, show α → β, from f) coe_injective + coe_one coe_mul coe_inv coe_div coe_pow coe_zpow + +@[to_additive] instance [comm_group β] : comm_group (α →ₛ β) := +function.injective.comm_group (λ f, show α → β, from f) coe_injective + coe_one coe_mul coe_inv coe_div coe_pow coe_zpow + +instance [semiring K] [add_comm_monoid β] [module K β] : module K (α →ₛ β) := +function.injective.module K ⟨λ f, show α → β, from f, coe_zero, coe_add⟩ + coe_injective coe_smul + +lemma smul_eq_map [has_smul K β] (k : K) (f : α →ₛ β) : k • f = f.map ((•) k) := rfl + +instance [preorder β] : preorder (α →ₛ β) := +{ le_refl := λf a, le_rfl, + le_trans := λf g h hfg hgh a, le_trans (hfg _) (hgh a), + .. simple_func.has_le } + +instance [partial_order β] : partial_order (α →ₛ β) := +{ le_antisymm := assume f g hfg hgf, ext $ assume a, le_antisymm (hfg a) (hgf a), + .. simple_func.preorder } + +instance [has_le β] [order_bot β] : order_bot (α →ₛ β) := +{ bot := const α ⊥, bot_le := λf a, bot_le } + +instance [has_le β] [order_top β] : order_top (α →ₛ β) := +{ top := const α ⊤, le_top := λf a, le_top } + +instance [semilattice_inf β] : semilattice_inf (α →ₛ β) := +{ inf := (⊓), + inf_le_left := assume f g a, inf_le_left, + inf_le_right := assume f g a, inf_le_right, + le_inf := assume f g h hfh hgh a, le_inf (hfh a) (hgh a), + .. simple_func.partial_order } + +instance [semilattice_sup β] : semilattice_sup (α →ₛ β) := +{ sup := (⊔), + le_sup_left := assume f g a, le_sup_left, + le_sup_right := assume f g a, le_sup_right, + sup_le := assume f g h hfh hgh a, sup_le (hfh a) (hgh a), + .. simple_func.partial_order } + +instance [lattice β] : lattice (α →ₛ β) := +{ .. simple_func.semilattice_sup,.. simple_func.semilattice_inf } + +instance [has_le β] [bounded_order β] : bounded_order (α →ₛ β) := +{ .. simple_func.order_bot, .. simple_func.order_top } + +lemma finset_sup_apply [semilattice_sup β] [order_bot β] {f : γ → α →ₛ β} (s : finset γ) (a : α) : + s.sup f a = s.sup (λc, f c a) := +begin + refine finset.induction_on s rfl _, + assume a s hs ih, + rw [finset.sup_insert, finset.sup_insert, sup_apply, ih] +end + +section restrict + +variables [has_zero β] + +/-- Restrict a simple function `f : α →ₛ β` to a set `s`. If `s` is measurable, +then `f.restrict s a = if a ∈ s then f a else 0`, otherwise `f.restrict s = const α 0`. -/ +def restrict (f : α →ₛ β) (s : set α) : α →ₛ β := +if hs : measurable_set s then piecewise s hs f 0 else 0 + +theorem restrict_of_not_measurable {f : α →ₛ β} {s : set α} + (hs : ¬measurable_set s) : + restrict f s = 0 := +dif_neg hs + +@[simp] theorem coe_restrict (f : α →ₛ β) {s : set α} (hs : measurable_set s) : + ⇑(restrict f s) = indicator s f := +by { rw [restrict, dif_pos hs], refl } + +@[simp] theorem restrict_univ (f : α →ₛ β) : restrict f univ = f := +by simp [restrict] + +@[simp] theorem restrict_empty (f : α →ₛ β) : restrict f ∅ = 0 := +by simp [restrict] + +theorem map_restrict_of_zero [has_zero γ] {g : β → γ} (hg : g 0 = 0) (f : α →ₛ β) (s : set α) : + (f.restrict s).map g = (f.map g).restrict s := +ext $ λ x, +if hs : measurable_set s then by simp [hs, set.indicator_comp_of_zero hg] +else by simp [restrict_of_not_measurable hs, hg] + +theorem map_coe_ennreal_restrict (f : α →ₛ ℝ≥0) (s : set α) : + (f.restrict s).map (coe : ℝ≥0 → ℝ≥0∞) = (f.map coe).restrict s := +map_restrict_of_zero ennreal.coe_zero _ _ + +theorem map_coe_nnreal_restrict (f : α →ₛ ℝ≥0) (s : set α) : + (f.restrict s).map (coe : ℝ≥0 → ℝ) = (f.map coe).restrict s := +map_restrict_of_zero nnreal.coe_zero _ _ + +theorem restrict_apply (f : α →ₛ β) {s : set α} (hs : measurable_set s) (a) : + restrict f s a = indicator s f a := +by simp only [f.coe_restrict hs] + +theorem restrict_preimage (f : α →ₛ β) {s : set α} (hs : measurable_set s) + {t : set β} (ht : (0:β) ∉ t) : restrict f s ⁻¹' t = s ∩ f ⁻¹' t := +by simp [hs, indicator_preimage_of_not_mem _ _ ht, inter_comm] + +theorem restrict_preimage_singleton (f : α →ₛ β) {s : set α} (hs : measurable_set s) + {r : β} (hr : r ≠ 0) : restrict f s ⁻¹' {r} = s ∩ f ⁻¹' {r} := +f.restrict_preimage hs hr.symm + +lemma mem_restrict_range {r : β} {s : set α} {f : α →ₛ β} (hs : measurable_set s) : + r ∈ (restrict f s).range ↔ (r = 0 ∧ s ≠ univ) ∨ (r ∈ f '' s) := +by rw [← finset.mem_coe, coe_range, coe_restrict _ hs, mem_range_indicator] + +lemma mem_image_of_mem_range_restrict {r : β} {s : set α} {f : α →ₛ β} + (hr : r ∈ (restrict f s).range) (h0 : r ≠ 0) : + r ∈ f '' s := +if hs : measurable_set s then by simpa [mem_restrict_range hs, h0] using hr +else by { rw [restrict_of_not_measurable hs] at hr, + exact (h0 $ eq_zero_of_mem_range_zero hr).elim } + +@[mono] lemma restrict_mono [preorder β] (s : set α) {f g : α →ₛ β} (H : f ≤ g) : + f.restrict s ≤ g.restrict s := +if hs : measurable_set s then λ x, by simp only [coe_restrict _ hs, indicator_le_indicator (H x)] +else by simp only [restrict_of_not_measurable hs, le_refl] + +end restrict + +section approx + +section +variables [semilattice_sup β] [order_bot β] [has_zero β] + +/-- Fix a sequence `i : ℕ → β`. Given a function `α → β`, its `n`-th approximation +by simple functions is defined so that in case `β = ℝ≥0∞` it sends each `a` to the supremum +of the set `{i k | k ≤ n ∧ i k ≤ f a}`, see `approx_apply` and `supr_approx_apply` for details. -/ +def approx (i : ℕ → β) (f : α → β) (n : ℕ) : α →ₛ β := +(finset.range n).sup (λk, restrict (const α (i k)) {a:α | i k ≤ f a}) + +lemma approx_apply [topological_space β] [order_closed_topology β] [measurable_space β] + [opens_measurable_space β] {i : ℕ → β} {f : α → β} {n : ℕ} (a : α) (hf : measurable f) : + (approx i f n : α →ₛ β) a = (finset.range n).sup (λk, if i k ≤ f a then i k else 0) := +begin + dsimp only [approx], + rw [finset_sup_apply], + congr, + funext k, + rw [restrict_apply], + refl, + exact (hf measurable_set_Ici) +end + +lemma monotone_approx (i : ℕ → β) (f : α → β) : monotone (approx i f) := +assume n m h, finset.sup_mono $ finset.range_subset.2 h + +lemma approx_comp [topological_space β] [order_closed_topology β] [measurable_space β] + [opens_measurable_space β] [measurable_space γ] + {i : ℕ → β} {f : γ → β} {g : α → γ} {n : ℕ} (a : α) + (hf : measurable f) (hg : measurable g) : + (approx i (f ∘ g) n : α →ₛ β) a = (approx i f n : γ →ₛ β) (g a) := +by rw [approx_apply _ hf, approx_apply _ (hf.comp hg)] + +end + +lemma supr_approx_apply [topological_space β] [complete_lattice β] [order_closed_topology β] + [has_zero β] [measurable_space β] [opens_measurable_space β] + (i : ℕ → β) (f : α → β) (a : α) (hf : measurable f) (h_zero : (0 : β) = ⊥) : + (⨆n, (approx i f n : α →ₛ β) a) = (⨆k (h : i k ≤ f a), i k) := +begin + refine le_antisymm (supr_le $ assume n, _) (supr_le $ assume k, supr_le $ assume hk, _), + { rw [approx_apply a hf, h_zero], + refine finset.sup_le (assume k hk, _), + split_ifs, + exact le_supr_of_le k (le_supr _ h), + exact bot_le }, + { refine le_supr_of_le (k+1) _, + rw [approx_apply a hf], + have : k ∈ finset.range (k+1) := finset.mem_range.2 (nat.lt_succ_self _), + refine le_trans (le_of_eq _) (finset.le_sup this), + rw [if_pos hk] } +end + +end approx + +section eapprox + +/-- A sequence of `ℝ≥0∞`s such that its range is the set of non-negative rational numbers. -/ +def ennreal_rat_embed (n : ℕ) : ℝ≥0∞ := +ennreal.of_real ((encodable.decode ℚ n).get_or_else (0 : ℚ)) + +lemma ennreal_rat_embed_encode (q : ℚ) : + ennreal_rat_embed (encodable.encode q) = real.to_nnreal q := +by rw [ennreal_rat_embed, encodable.encodek]; refl + +/-- Approximate a function `α → ℝ≥0∞` by a sequence of simple functions. -/ +def eapprox : (α → ℝ≥0∞) → ℕ → α →ₛ ℝ≥0∞ := +approx ennreal_rat_embed + +lemma eapprox_lt_top (f : α → ℝ≥0∞) (n : ℕ) (a : α) : eapprox f n a < ∞ := +begin + simp only [eapprox, approx, finset_sup_apply, finset.sup_lt_iff, with_top.zero_lt_top, + finset.mem_range, ennreal.bot_eq_zero, restrict], + assume b hb, + split_ifs, + { simp only [coe_zero, coe_piecewise, piecewise_eq_indicator, coe_const], + calc {a : α | ennreal_rat_embed b ≤ f a}.indicator (λ x, ennreal_rat_embed b) a + ≤ ennreal_rat_embed b : indicator_le_self _ _ a + ... < ⊤ : ennreal.coe_lt_top }, + { exact with_top.zero_lt_top }, +end + +@[mono] lemma monotone_eapprox (f : α → ℝ≥0∞) : monotone (eapprox f) := +monotone_approx _ f + +lemma supr_eapprox_apply (f : α → ℝ≥0∞) (hf : measurable f) (a : α) : + (⨆n, (eapprox f n : α →ₛ ℝ≥0∞) a) = f a := +begin + rw [eapprox, supr_approx_apply ennreal_rat_embed f a hf rfl], + refine le_antisymm (supr_le $ assume i, supr_le $ assume hi, hi) (le_of_not_gt _), + assume h, + rcases ennreal.lt_iff_exists_rat_btwn.1 h with ⟨q, hq, lt_q, q_lt⟩, + have : (real.to_nnreal q : ℝ≥0∞) ≤ + (⨆ (k : ℕ) (h : ennreal_rat_embed k ≤ f a), ennreal_rat_embed k), + { refine le_supr_of_le (encodable.encode q) _, + rw [ennreal_rat_embed_encode q], + refine le_supr_of_le (le_of_lt q_lt) _, + exact le_rfl }, + exact lt_irrefl _ (lt_of_le_of_lt this lt_q) +end + +lemma eapprox_comp [measurable_space γ] {f : γ → ℝ≥0∞} {g : α → γ} {n : ℕ} + (hf : measurable f) (hg : measurable g) : + (eapprox (f ∘ g) n : α → ℝ≥0∞) = (eapprox f n : γ →ₛ ℝ≥0∞) ∘ g := +funext $ assume a, approx_comp a hf hg + +/-- Approximate a function `α → ℝ≥0∞` by a series of simple functions taking their values +in `ℝ≥0`. -/ +def eapprox_diff (f : α → ℝ≥0∞) : ∀ (n : ℕ), α →ₛ ℝ≥0 +| 0 := (eapprox f 0).map ennreal.to_nnreal +| (n+1) := (eapprox f (n+1) - eapprox f n).map ennreal.to_nnreal + +lemma sum_eapprox_diff (f : α → ℝ≥0∞) (n : ℕ) (a : α) : + (∑ k in finset.range (n+1), (eapprox_diff f k a : ℝ≥0∞)) = eapprox f n a := +begin + induction n with n IH, + { simp only [nat.nat_zero_eq_zero, finset.sum_singleton, finset.range_one], refl }, + { rw [finset.sum_range_succ, nat.succ_eq_add_one, IH, eapprox_diff, coe_map, function.comp_app, + coe_sub, pi.sub_apply, ennreal.coe_to_nnreal, + add_tsub_cancel_of_le (monotone_eapprox f (nat.le_succ _) _)], + apply (lt_of_le_of_lt _ (eapprox_lt_top f (n+1) a)).ne, + rw tsub_le_iff_right, + exact le_self_add }, +end + +lemma tsum_eapprox_diff (f : α → ℝ≥0∞) (hf : measurable f) (a : α) : + (∑' n, (eapprox_diff f n a : ℝ≥0∞)) = f a := +by simp_rw [ennreal.tsum_eq_supr_nat' (tendsto_add_at_top_nat 1), sum_eapprox_diff, + supr_eapprox_apply f hf a] + +end eapprox + +end measurable + +section measure +variables {m : measurable_space α} {μ ν : measure α} + +/-- Integral of a simple function whose codomain is `ℝ≥0∞`. -/ +def lintegral {m : measurable_space α} (f : α →ₛ ℝ≥0∞) (μ : measure α) : ℝ≥0∞ := +∑ x in f.range, x * μ (f ⁻¹' {x}) + +lemma lintegral_eq_of_subset (f : α →ₛ ℝ≥0∞) {s : finset ℝ≥0∞} + (hs : ∀ x, f x ≠ 0 → μ (f ⁻¹' {f x}) ≠ 0 → f x ∈ s) : + f.lintegral μ = ∑ x in s, x * μ (f ⁻¹' {x}) := +begin + refine finset.sum_bij_ne_zero (λr _ _, r) _ _ _ _, + { simpa only [forall_range_iff, mul_ne_zero_iff, and_imp] }, + { intros, assumption }, + { intros b _ hb, + refine ⟨b, _, hb, rfl⟩, + rw [mem_range, ← preimage_singleton_nonempty], + exact nonempty_of_measure_ne_zero (mul_ne_zero_iff.1 hb).2 }, + { intros, refl } +end + +lemma lintegral_eq_of_subset' (f : α →ₛ ℝ≥0∞) {s : finset ℝ≥0∞} + (hs : f.range \ {0} ⊆ s) : + f.lintegral μ = ∑ x in s, x * μ (f ⁻¹' {x}) := +f.lintegral_eq_of_subset $ λ x hfx _, hs $ + finset.mem_sdiff.2 ⟨f.mem_range_self x, mt finset.mem_singleton.1 hfx⟩ + +/-- Calculate the integral of `(g ∘ f)`, where `g : β → ℝ≥0∞` and `f : α →ₛ β`. -/ +lemma map_lintegral (g : β → ℝ≥0∞) (f : α →ₛ β) : + (f.map g).lintegral μ = ∑ x in f.range, g x * μ (f ⁻¹' {x}) := +begin + simp only [lintegral, range_map], + refine finset.sum_image' _ (assume b hb, _), + rcases mem_range.1 hb with ⟨a, rfl⟩, + rw [map_preimage_singleton, ← f.sum_measure_preimage_singleton, finset.mul_sum], + refine finset.sum_congr _ _, + { congr }, + { assume x, simp only [finset.mem_filter], rintro ⟨_, h⟩, rw h }, +end + +lemma add_lintegral (f g : α →ₛ ℝ≥0∞) : (f + g).lintegral μ = f.lintegral μ + g.lintegral μ := +calc (f + g).lintegral μ = + ∑ x in (pair f g).range, (x.1 * μ (pair f g ⁻¹' {x}) + x.2 * μ (pair f g ⁻¹' {x})) : + by rw [add_eq_map₂, map_lintegral]; exact finset.sum_congr rfl (assume a ha, add_mul _ _ _) + ... = ∑ x in (pair f g).range, x.1 * μ (pair f g ⁻¹' {x}) + + ∑ x in (pair f g).range, x.2 * μ (pair f g ⁻¹' {x}) : by rw [finset.sum_add_distrib] + ... = ((pair f g).map prod.fst).lintegral μ + ((pair f g).map prod.snd).lintegral μ : + by rw [map_lintegral, map_lintegral] + ... = lintegral f μ + lintegral g μ : rfl + +lemma const_mul_lintegral (f : α →ₛ ℝ≥0∞) (x : ℝ≥0∞) : + (const α x * f).lintegral μ = x * f.lintegral μ := +calc (f.map (λa, x * a)).lintegral μ = ∑ r in f.range, x * r * μ (f ⁻¹' {r}) : + map_lintegral _ _ + ... = ∑ r in f.range, x * (r * μ (f ⁻¹' {r})) : + finset.sum_congr rfl (assume a ha, mul_assoc _ _ _) + ... = x * f.lintegral μ : + finset.mul_sum.symm + +/-- Integral of a simple function `α →ₛ ℝ≥0∞` as a bilinear map. -/ +def lintegralₗ {m : measurable_space α} : (α →ₛ ℝ≥0∞) →ₗ[ℝ≥0∞] measure α →ₗ[ℝ≥0∞] ℝ≥0∞ := +{ to_fun := λ f, + { to_fun := lintegral f, + map_add' := by simp [lintegral, mul_add, finset.sum_add_distrib], + map_smul' := λ c μ, by simp [lintegral, mul_left_comm _ c, finset.mul_sum] }, + map_add' := λ f g, linear_map.ext (λ μ, add_lintegral f g), + map_smul' := λ c f, linear_map.ext (λ μ, const_mul_lintegral f c) } + +@[simp] lemma zero_lintegral : (0 : α →ₛ ℝ≥0∞).lintegral μ = 0 := +linear_map.ext_iff.1 lintegralₗ.map_zero μ + +lemma lintegral_add {ν} (f : α →ₛ ℝ≥0∞) : f.lintegral (μ + ν) = f.lintegral μ + f.lintegral ν := +(lintegralₗ f).map_add μ ν + +lemma lintegral_smul (f : α →ₛ ℝ≥0∞) (c : ℝ≥0∞) : + f.lintegral (c • μ) = c • f.lintegral μ := +(lintegralₗ f).map_smul c μ + +@[simp] lemma lintegral_zero [measurable_space α] (f : α →ₛ ℝ≥0∞) : + f.lintegral 0 = 0 := +(lintegralₗ f).map_zero + +lemma lintegral_sum {m : measurable_space α} {ι} (f : α →ₛ ℝ≥0∞) (μ : ι → measure α) : + f.lintegral (measure.sum μ) = ∑' i, f.lintegral (μ i) := +begin + simp only [lintegral, measure.sum_apply, f.measurable_set_preimage, ← finset.tsum_subtype, + ← ennreal.tsum_mul_left], + apply ennreal.tsum_comm +end + +lemma restrict_lintegral (f : α →ₛ ℝ≥0∞) {s : set α} (hs : measurable_set s) : + (restrict f s).lintegral μ = ∑ r in f.range, r * μ (f ⁻¹' {r} ∩ s) := +calc (restrict f s).lintegral μ = ∑ r in f.range, r * μ (restrict f s ⁻¹' {r}) : + lintegral_eq_of_subset _ $ λ x hx, if hxs : x ∈ s + then λ _, by simp only [f.restrict_apply hs, indicator_of_mem hxs, mem_range_self] + else false.elim $ hx $ by simp [*] +... = ∑ r in f.range, r * μ (f ⁻¹' {r} ∩ s) : + finset.sum_congr rfl $ forall_range_iff.2 $ λ b, if hb : f b = 0 then by simp only [hb, zero_mul] + else by rw [restrict_preimage_singleton _ hs hb, inter_comm] + +lemma lintegral_restrict {m : measurable_space α} (f : α →ₛ ℝ≥0∞) (s : set α) (μ : measure α) : + f.lintegral (μ.restrict s) = ∑ y in f.range, y * μ (f ⁻¹' {y} ∩ s) := +by simp only [lintegral, measure.restrict_apply, f.measurable_set_preimage] + +lemma restrict_lintegral_eq_lintegral_restrict (f : α →ₛ ℝ≥0∞) {s : set α} + (hs : measurable_set s) : + (restrict f s).lintegral μ = f.lintegral (μ.restrict s) := +by rw [f.restrict_lintegral hs, lintegral_restrict] + +lemma const_lintegral (c : ℝ≥0∞) : (const α c).lintegral μ = c * μ univ := +begin + rw [lintegral], + casesI is_empty_or_nonempty α, + { simp [μ.eq_zero_of_is_empty] }, + { simp [preimage_const_of_mem] }, +end + +lemma const_lintegral_restrict (c : ℝ≥0∞) (s : set α) : + (const α c).lintegral (μ.restrict s) = c * μ s := +by rw [const_lintegral, measure.restrict_apply measurable_set.univ, univ_inter] + +lemma restrict_const_lintegral (c : ℝ≥0∞) {s : set α} (hs : measurable_set s) : + ((const α c).restrict s).lintegral μ = c * μ s := +by rw [restrict_lintegral_eq_lintegral_restrict _ hs, const_lintegral_restrict] + +lemma le_sup_lintegral (f g : α →ₛ ℝ≥0∞) : f.lintegral μ ⊔ g.lintegral μ ≤ (f ⊔ g).lintegral μ := +calc f.lintegral μ ⊔ g.lintegral μ = + ((pair f g).map prod.fst).lintegral μ ⊔ ((pair f g).map prod.snd).lintegral μ : rfl + ... ≤ ∑ x in (pair f g).range, (x.1 ⊔ x.2) * μ (pair f g ⁻¹' {x}) : + begin + rw [map_lintegral, map_lintegral], + refine sup_le _ _; + refine finset.sum_le_sum (λ a _, mul_le_mul_right' _ _), + exact le_sup_left, + exact le_sup_right + end + ... = (f ⊔ g).lintegral μ : by rw [sup_eq_map₂, map_lintegral] + +/-- `simple_func.lintegral` is monotone both in function and in measure. -/ +@[mono] lemma lintegral_mono {f g : α →ₛ ℝ≥0∞} (hfg : f ≤ g) (hμν : μ ≤ ν) : + f.lintegral μ ≤ g.lintegral ν := +calc f.lintegral μ ≤ f.lintegral μ ⊔ g.lintegral μ : le_sup_left + ... ≤ (f ⊔ g).lintegral μ : le_sup_lintegral _ _ + ... = g.lintegral μ : by rw [sup_of_le_right hfg] + ... ≤ g.lintegral ν : finset.sum_le_sum $ λ y hy, ennreal.mul_left_mono $ + hμν _ (g.measurable_set_preimage _) + +/-- `simple_func.lintegral` depends only on the measures of `f ⁻¹' {y}`. -/ +lemma lintegral_eq_of_measure_preimage [measurable_space β] {f : α →ₛ ℝ≥0∞} {g : β →ₛ ℝ≥0∞} + {ν : measure β} (H : ∀ y, μ (f ⁻¹' {y}) = ν (g ⁻¹' {y})) : + f.lintegral μ = g.lintegral ν := +begin + simp only [lintegral, ← H], + apply lintegral_eq_of_subset, + simp only [H], + intros, + exact mem_range_of_measure_ne_zero ‹_› +end + +/-- If two simple functions are equal a.e., then their `lintegral`s are equal. -/ +lemma lintegral_congr {f g : α →ₛ ℝ≥0∞} (h : f =ᵐ[μ] g) : + f.lintegral μ = g.lintegral μ := +lintegral_eq_of_measure_preimage $ λ y, measure_congr $ + eventually.set_eq $ h.mono $ λ x hx, by simp [hx] + +lemma lintegral_map' {β} [measurable_space β] {μ' : measure β} (f : α →ₛ ℝ≥0∞) (g : β →ₛ ℝ≥0∞) + (m' : α → β) (eq : ∀ a, f a = g (m' a)) (h : ∀s, measurable_set s → μ' s = μ (m' ⁻¹' s)) : + f.lintegral μ = g.lintegral μ' := +lintegral_eq_of_measure_preimage $ λ y, +by { simp only [preimage, eq], exact (h (g ⁻¹' {y}) (g.measurable_set_preimage _)).symm } + +lemma lintegral_map {β} [measurable_space β] (g : β →ₛ ℝ≥0∞) {f : α → β} (hf : measurable f) : + g.lintegral (measure.map f μ) = (g.comp f hf).lintegral μ := +eq.symm $ lintegral_map' _ _ f (λ a, rfl) (λ s hs, measure.map_apply hf hs) + +end measure + +section fin_meas_supp + +open finset function + +lemma support_eq [measurable_space α] [has_zero β] (f : α →ₛ β) : + support f = ⋃ y ∈ f.range.filter (λ y, y ≠ 0), f ⁻¹' {y} := +set.ext $ λ x, by simp only [mem_support, set.mem_preimage, mem_filter, mem_range_self, true_and, + exists_prop, mem_Union, set.mem_range, mem_singleton_iff, exists_eq_right'] + +variables {m : measurable_space α} [has_zero β] [has_zero γ] {μ : measure α} {f : α →ₛ β} + +lemma measurable_set_support [measurable_space α] (f : α →ₛ β) : measurable_set (support f) := +by { rw f.support_eq, exact finset.measurable_set_bUnion _ (λ y hy, measurable_set_fiber _ _), } + +/-- A `simple_func` has finite measure support if it is equal to `0` outside of a set of finite +measure. -/ +protected def fin_meas_supp {m : measurable_space α} (f : α →ₛ β) (μ : measure α) : Prop := +f =ᶠ[μ.cofinite] 0 + +lemma fin_meas_supp_iff_support : f.fin_meas_supp μ ↔ μ (support f) < ∞ := iff.rfl + +lemma fin_meas_supp_iff : f.fin_meas_supp μ ↔ ∀ y ≠ 0, μ (f ⁻¹' {y}) < ∞ := +begin + split, + { refine λ h y hy, lt_of_le_of_lt (measure_mono _) h, + exact λ x hx (H : f x = 0), hy $ H ▸ eq.symm hx }, + { intro H, + rw [fin_meas_supp_iff_support, support_eq], + refine lt_of_le_of_lt (measure_bUnion_finset_le _ _) (sum_lt_top _), + exact λ y hy, (H y (finset.mem_filter.1 hy).2).ne } +end + +namespace fin_meas_supp + +lemma meas_preimage_singleton_ne_zero (h : f.fin_meas_supp μ) {y : β} (hy : y ≠ 0) : + μ (f ⁻¹' {y}) < ∞ := +fin_meas_supp_iff.1 h y hy + +protected lemma map {g : β → γ} (hf : f.fin_meas_supp μ) (hg : g 0 = 0) : + (f.map g).fin_meas_supp μ := +flip lt_of_le_of_lt hf (measure_mono $ support_comp_subset hg f) + +lemma of_map {g : β → γ} (h : (f.map g).fin_meas_supp μ) (hg : ∀b, g b = 0 → b = 0) : + f.fin_meas_supp μ := +flip lt_of_le_of_lt h $ measure_mono $ support_subset_comp hg _ + +lemma map_iff {g : β → γ} (hg : ∀ {b}, g b = 0 ↔ b = 0) : + (f.map g).fin_meas_supp μ ↔ f.fin_meas_supp μ := +⟨λ h, h.of_map $ λ b, hg.1, λ h, h.map $ hg.2 rfl⟩ + +protected lemma pair {g : α →ₛ γ} (hf : f.fin_meas_supp μ) (hg : g.fin_meas_supp μ) : + (pair f g).fin_meas_supp μ := +calc μ (support $ pair f g) = μ (support f ∪ support g) : congr_arg μ $ support_prod_mk f g +... ≤ μ (support f) + μ (support g) : measure_union_le _ _ +... < _ : add_lt_top.2 ⟨hf, hg⟩ + +protected lemma map₂ [has_zero δ] (hf : f.fin_meas_supp μ) + {g : α →ₛ γ} (hg : g.fin_meas_supp μ) {op : β → γ → δ} (H : op 0 0 = 0) : + ((pair f g).map (function.uncurry op)).fin_meas_supp μ := +(hf.pair hg).map H + +protected lemma add {β} [add_monoid β] {f g : α →ₛ β} (hf : f.fin_meas_supp μ) + (hg : g.fin_meas_supp μ) : + (f + g).fin_meas_supp μ := +by { rw [add_eq_map₂], exact hf.map₂ hg (zero_add 0) } + +protected lemma mul {β} [monoid_with_zero β] {f g : α →ₛ β} (hf : f.fin_meas_supp μ) + (hg : g.fin_meas_supp μ) : + (f * g).fin_meas_supp μ := +by { rw [mul_eq_map₂], exact hf.map₂ hg (zero_mul 0) } + +lemma lintegral_lt_top {f : α →ₛ ℝ≥0∞} (hm : f.fin_meas_supp μ) (hf : ∀ᵐ a ∂μ, f a ≠ ∞) : + f.lintegral μ < ∞ := +begin + refine sum_lt_top (λ a ha, _), + rcases eq_or_ne a ∞ with rfl|ha, + { simp only [ae_iff, ne.def, not_not] at hf, + simp [set.preimage, hf] }, + { by_cases ha0 : a = 0, + { subst a, rwa [zero_mul] }, + { exact mul_ne_top ha (fin_meas_supp_iff.1 hm _ ha0).ne } } +end + +lemma of_lintegral_ne_top {f : α →ₛ ℝ≥0∞} (h : f.lintegral μ ≠ ∞) : f.fin_meas_supp μ := +begin + refine fin_meas_supp_iff.2 (λ b hb, _), + rw [f.lintegral_eq_of_subset' (finset.subset_insert b _)] at h, + refine ennreal.lt_top_of_mul_ne_top_right _ hb, + exact (lt_top_of_sum_ne_top h (finset.mem_insert_self _ _)).ne +end + +lemma iff_lintegral_lt_top {f : α →ₛ ℝ≥0∞} (hf : ∀ᵐ a ∂μ, f a ≠ ∞) : + f.fin_meas_supp μ ↔ f.lintegral μ < ∞ := +⟨λ h, h.lintegral_lt_top hf, λ h, of_lintegral_ne_top h.ne⟩ + +end fin_meas_supp + +end fin_meas_supp + +/-- To prove something for an arbitrary simple function, it suffices to show +that the property holds for (multiples of) characteristic functions and is closed under +addition (of functions with disjoint support). + +It is possible to make the hypotheses in `h_add` a bit stronger, and such conditions can be added +once we need them (for example it is only necessary to consider the case where `g` is a multiple +of a characteristic function, and that this multiple doesn't appear in the image of `f`) -/ +@[elab_as_eliminator] +protected lemma induction {α γ} [measurable_space α] [add_monoid γ] {P : simple_func α γ → Prop} + (h_ind : ∀ c {s} (hs : measurable_set s), + P (simple_func.piecewise s hs (simple_func.const _ c) (simple_func.const _ 0))) + (h_add : ∀ ⦃f g : simple_func α γ⦄, disjoint (support f) (support g) → P f → P g → P (f + g)) + (f : simple_func α γ) : P f := +begin + generalize' h : f.range \ {0} = s, + rw [← finset.coe_inj, finset.coe_sdiff, finset.coe_singleton, simple_func.coe_range] at h, + revert s f h, refine finset.induction _ _, + { intros f hf, rw [finset.coe_empty, diff_eq_empty, range_subset_singleton] at hf, + convert h_ind 0 measurable_set.univ, ext x, simp [hf] }, + { intros x s hxs ih f hf, + have mx := f.measurable_set_preimage {x}, + let g := simple_func.piecewise (f ⁻¹' {x}) mx 0 f, + have Pg : P g, + { apply ih, simp only [g, simple_func.coe_piecewise, range_piecewise], + rw [image_compl_preimage, union_diff_distrib, diff_diff_comm, hf, finset.coe_insert, + insert_diff_self_of_not_mem, diff_eq_empty.mpr, set.empty_union], + { rw [set.image_subset_iff], convert set.subset_univ _, + exact preimage_const_of_mem (mem_singleton _) }, + { rwa [finset.mem_coe] }}, + convert h_add _ Pg (h_ind x mx), + { ext1 y, by_cases hy : y ∈ f ⁻¹' {x}; [simpa [hy], simp [hy]] }, + rw disjoint_iff_inf_le, + rintro y, by_cases hy : y ∈ f ⁻¹' {x}; simp [hy] } +end + +end simple_func + +end measure_theory + +open measure_theory measure_theory.simple_func +/-- To prove something for an arbitrary measurable function into `ℝ≥0∞`, it suffices to show +that the property holds for (multiples of) characteristic functions and is closed under addition +and supremum of increasing sequences of functions. + +It is possible to make the hypotheses in the induction steps a bit stronger, and such conditions +can be added once we need them (for example in `h_add` it is only necessary to consider the sum of +a simple function with a multiple of a characteristic function and that the intersection +of their images is a subset of `{0}`. -/ +@[elab_as_eliminator] +theorem measurable.ennreal_induction {α} [measurable_space α] {P : (α → ℝ≥0∞) → Prop} + (h_ind : ∀ (c : ℝ≥0∞) ⦃s⦄, measurable_set s → P (indicator s (λ _, c))) + (h_add : ∀ ⦃f g : α → ℝ≥0∞⦄, disjoint (support f) (support g) → measurable f → measurable g → + P f → P g → P (f + g)) + (h_supr : ∀ ⦃f : ℕ → α → ℝ≥0∞⦄ (hf : ∀n, measurable (f n)) (h_mono : monotone f) + (hP : ∀ n, P (f n)), P (λ x, ⨆ n, f n x)) + ⦃f : α → ℝ≥0∞⦄ (hf : measurable f) : P f := +begin + convert h_supr (λ n, (eapprox f n).measurable) (monotone_eapprox f) _, + { ext1 x, rw [supr_eapprox_apply f hf] }, + { exact λ n, simple_func.induction (λ c s hs, h_ind c hs) + (λ f g hfg hf hg, h_add hfg f.measurable g.measurable hf hg) (eapprox f n) } +end diff --git a/src/measure_theory/function/simple_func_dense.lean b/src/measure_theory/function/simple_func_dense.lean index 5c7ca36511d8f..b9d7038805441 100644 --- a/src/measure_theory/function/simple_func_dense.lean +++ b/src/measure_theory/function/simple_func_dense.lean @@ -3,7 +3,7 @@ Copyright (c) 2019 Zhouhang Zhou. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Zhouhang Zhou, Yury Kudryashov, Heather Macbeth -/ -import measure_theory.integral.lebesgue +import measure_theory.function.simple_func /-! # Density of simple functions diff --git a/src/measure_theory/function/strongly_measurable/basic.lean b/src/measure_theory/function/strongly_measurable/basic.lean index 425643486a0f9..610bde59e6b99 100644 --- a/src/measure_theory/function/strongly_measurable/basic.lean +++ b/src/measure_theory/function/strongly_measurable/basic.lean @@ -5,6 +5,7 @@ Authors: Rémy Degenne, Sébastien Gouëzel -/ import analysis.normed_space.bounded_linear_maps import topology.metric_space.metrizable +import measure_theory.integral.lebesgue import measure_theory.function.simple_func_dense /-! diff --git a/src/measure_theory/integral/lebesgue.lean b/src/measure_theory/integral/lebesgue.lean index 60c905b776f42..685b546bb933e 100644 --- a/src/measure_theory/integral/lebesgue.lean +++ b/src/measure_theory/integral/lebesgue.lean @@ -3,22 +3,14 @@ Copyright (c) 2018 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Johannes Hölzl -/ -import measure_theory.measure.mutually_singular -import measure_theory.constructions.borel_space -import algebra.indicator_function -import algebra.support import dynamics.ergodic.measure_preserving +import measure_theory.function.simple_func +import measure_theory.measure.mutually_singular /-! -# Lebesgue integral for `ℝ≥0∞`-valued functions - -We define simple functions and show that each Borel measurable function on `ℝ≥0∞` can be -approximated by a sequence of simple functions. +# Lower Lebesgue integral for `ℝ≥0∞`-valued functions -To prove something for an arbitrary measurable function into `ℝ≥0∞`, the theorem -`measurable.ennreal_induction` shows that is it sufficient to show that the property holds for -(multiples of) characteristic functions and is closed under addition and supremum of increasing -sequences of functions. +We define the lower Lebesgue integral of an `ℝ≥0∞`-valued function. ## Notation @@ -40,967 +32,14 @@ open_locale classical topology big_operators nnreal ennreal measure_theory namespace measure_theory -variables {α β γ δ : Type*} - -/-- A function `f` from a measurable space to any type is called *simple*, -if every preimage `f ⁻¹' {x}` is measurable, and the range is finite. This structure bundles -a function with these properties. -/ -structure {u v} simple_func (α : Type u) [measurable_space α] (β : Type v) := -(to_fun : α → β) -(measurable_set_fiber' : ∀ x, measurable_set (to_fun ⁻¹' {x})) -(finite_range' : (set.range to_fun).finite) - local infixr ` →ₛ `:25 := simple_func -namespace simple_func - -section measurable -variables [measurable_space α] -instance has_coe_to_fun : has_coe_to_fun (α →ₛ β) (λ _, α → β) := ⟨to_fun⟩ - -lemma coe_injective ⦃f g : α →ₛ β⦄ (H : (f : α → β) = g) : f = g := -by cases f; cases g; congr; exact H - -@[ext] theorem ext {f g : α →ₛ β} (H : ∀ a, f a = g a) : f = g := -coe_injective $ funext H - -lemma finite_range (f : α →ₛ β) : (set.range f).finite := f.finite_range' - -lemma measurable_set_fiber (f : α →ₛ β) (x : β) : measurable_set (f ⁻¹' {x}) := -f.measurable_set_fiber' x - -@[simp] lemma apply_mk (f : α → β) (h h') (x : α) : simple_func.mk f h h' x = f x := rfl - -/-- Simple function defined on the empty type. -/ -def of_is_empty [is_empty α] : α →ₛ β := -{ to_fun := is_empty_elim, - measurable_set_fiber' := λ x, subsingleton.measurable_set, - finite_range' := by simp [range_eq_empty] } - -/-- Range of a simple function `α →ₛ β` as a `finset β`. -/ -protected def range (f : α →ₛ β) : finset β := f.finite_range.to_finset - -@[simp] theorem mem_range {f : α →ₛ β} {b} : b ∈ f.range ↔ b ∈ range f := -finite.mem_to_finset _ - -theorem mem_range_self (f : α →ₛ β) (x : α) : f x ∈ f.range := mem_range.2 ⟨x, rfl⟩ - -@[simp] lemma coe_range (f : α →ₛ β) : (↑f.range : set β) = set.range f := -f.finite_range.coe_to_finset - -theorem mem_range_of_measure_ne_zero {f : α →ₛ β} {x : β} {μ : measure α} (H : μ (f ⁻¹' {x}) ≠ 0) : - x ∈ f.range := -let ⟨a, ha⟩ := nonempty_of_measure_ne_zero H in -mem_range.2 ⟨a, ha⟩ - -lemma forall_range_iff {f : α →ₛ β} {p : β → Prop} : - (∀ y ∈ f.range, p y) ↔ ∀ x, p (f x) := -by simp only [mem_range, set.forall_range_iff] - -lemma exists_range_iff {f : α →ₛ β} {p : β → Prop} : - (∃ y ∈ f.range, p y) ↔ ∃ x, p (f x) := -by simpa only [mem_range, exists_prop] using set.exists_range_iff - -lemma preimage_eq_empty_iff (f : α →ₛ β) (b : β) : f ⁻¹' {b} = ∅ ↔ b ∉ f.range := -preimage_singleton_eq_empty.trans $ not_congr mem_range.symm - -lemma exists_forall_le [nonempty β] [preorder β] [is_directed β (≤)] (f : α →ₛ β) : - ∃ C, ∀ x, f x ≤ C := -f.range.exists_le.imp $ λ C, forall_range_iff.1 - -/-- Constant function as a `simple_func`. -/ -def const (α) {β} [measurable_space α] (b : β) : α →ₛ β := -⟨λ a, b, λ x, measurable_set.const _, finite_range_const⟩ - -instance [inhabited β] : inhabited (α →ₛ β) := ⟨const _ default⟩ - -theorem const_apply (a : α) (b : β) : (const α b) a = b := rfl - -@[simp] theorem coe_const (b : β) : ⇑(const α b) = function.const α b := rfl - -@[simp] lemma range_const (α) [measurable_space α] [nonempty α] (b : β) : - (const α b).range = {b} := -finset.coe_injective $ by simp - -lemma range_const_subset (α) [measurable_space α] (b : β) : - (const α b).range ⊆ {b} := -finset.coe_subset.1 $ by simp - -lemma simple_func_bot {α} (f : @simple_func α ⊥ β) [nonempty β] : ∃ c, ∀ x, f x = c := -begin - have hf_meas := @simple_func.measurable_set_fiber α _ ⊥ f, - simp_rw measurable_space.measurable_set_bot_iff at hf_meas, - casesI is_empty_or_nonempty α, - { simp only [is_empty.forall_iff, exists_const], }, - { specialize hf_meas (f h.some), - cases hf_meas, - { exfalso, - refine set.not_mem_empty h.some _, - rw [← hf_meas, set.mem_preimage], - exact set.mem_singleton _, }, - { refine ⟨f h.some, λ x, _⟩, - have : x ∈ f ⁻¹' {f h.some}, - { rw hf_meas, exact set.mem_univ x, }, - rwa [set.mem_preimage, set.mem_singleton_iff] at this, }, }, -end - -lemma simple_func_bot' {α} [nonempty β] (f : @simple_func α ⊥ β) : - ∃ c, f = @simple_func.const α _ ⊥ c := -begin - obtain ⟨c, h_eq⟩ := simple_func_bot f, - refine ⟨c, _⟩, - ext1 x, - rw [h_eq x, simple_func.coe_const], -end - -lemma measurable_set_cut (r : α → β → Prop) (f : α →ₛ β) - (h : ∀b, measurable_set {a | r a b}) : measurable_set {a | r a (f a)} := -begin - have : {a | r a (f a)} = ⋃ b ∈ range f, {a | r a b} ∩ f ⁻¹' {b}, - { ext a, - suffices : r a (f a) ↔ ∃ i, r a (f i) ∧ f a = f i, by simpa, - exact ⟨λ h, ⟨a, ⟨h, rfl⟩⟩, λ ⟨a', ⟨h', e⟩⟩, e.symm ▸ h'⟩ }, - rw this, - exact measurable_set.bUnion f.finite_range.countable - (λ b _, measurable_set.inter (h b) (f.measurable_set_fiber _)) -end - -@[measurability] -theorem measurable_set_preimage (f : α →ₛ β) (s) : measurable_set (f ⁻¹' s) := -measurable_set_cut (λ _ b, b ∈ s) f (λ b, measurable_set.const (b ∈ s)) - -/-- A simple function is measurable -/ -@[measurability] -protected theorem measurable [measurable_space β] (f : α →ₛ β) : measurable f := -λ s _, measurable_set_preimage f s - -@[measurability] -protected theorem ae_measurable [measurable_space β] {μ : measure α} (f : α →ₛ β) : - ae_measurable f μ := -f.measurable.ae_measurable - -protected lemma sum_measure_preimage_singleton (f : α →ₛ β) {μ : measure α} (s : finset β) : - ∑ y in s, μ (f ⁻¹' {y}) = μ (f ⁻¹' ↑s) := -sum_measure_preimage_singleton _ (λ _ _, f.measurable_set_fiber _) - -lemma sum_range_measure_preimage_singleton (f : α →ₛ β) (μ : measure α) : - ∑ y in f.range, μ (f ⁻¹' {y}) = μ univ := -by rw [f.sum_measure_preimage_singleton, coe_range, preimage_range] - -/-- If-then-else as a `simple_func`. -/ -def piecewise (s : set α) (hs : measurable_set s) (f g : α →ₛ β) : α →ₛ β := -⟨s.piecewise f g, - λ x, by letI : measurable_space β := ⊤; exact - f.measurable.piecewise hs g.measurable trivial, - (f.finite_range.union g.finite_range).subset range_ite_subset⟩ - -@[simp] theorem coe_piecewise {s : set α} (hs : measurable_set s) (f g : α →ₛ β) : - ⇑(piecewise s hs f g) = s.piecewise f g := -rfl - -theorem piecewise_apply {s : set α} (hs : measurable_set s) (f g : α →ₛ β) (a) : - piecewise s hs f g a = if a ∈ s then f a else g a := -rfl - -@[simp] lemma piecewise_compl {s : set α} (hs : measurable_set sᶜ) (f g : α →ₛ β) : - piecewise sᶜ hs f g = piecewise s hs.of_compl g f := -coe_injective $ by simp [hs] - -@[simp] lemma piecewise_univ (f g : α →ₛ β) : piecewise univ measurable_set.univ f g = f := -coe_injective $ by simp - -@[simp] lemma piecewise_empty (f g : α →ₛ β) : piecewise ∅ measurable_set.empty f g = g := -coe_injective $ by simp - -lemma support_indicator [has_zero β] {s : set α} (hs : measurable_set s) (f : α →ₛ β) : - function.support (f.piecewise s hs (simple_func.const α 0)) = s ∩ function.support f := -set.support_indicator - -lemma range_indicator {s : set α} (hs : measurable_set s) - (hs_nonempty : s.nonempty) (hs_ne_univ : s ≠ univ) (x y : β) : - (piecewise s hs (const α x) (const α y)).range = {x, y} := -by simp only [← finset.coe_inj, coe_range, coe_piecewise, range_piecewise, coe_const, - finset.coe_insert, finset.coe_singleton, hs_nonempty.image_const, - (nonempty_compl.2 hs_ne_univ).image_const, singleton_union] - -lemma measurable_bind [measurable_space γ] (f : α →ₛ β) (g : β → α → γ) - (hg : ∀ b, measurable (g b)) : measurable (λ a, g (f a) a) := -λ s hs, f.measurable_set_cut (λ a b, g b a ∈ s) $ λ b, hg b hs - -/-- If `f : α →ₛ β` is a simple function and `g : β → α →ₛ γ` is a family of simple functions, -then `f.bind g` binds the first argument of `g` to `f`. In other words, `f.bind g a = g (f a) a`. -/ -def bind (f : α →ₛ β) (g : β → α →ₛ γ) : α →ₛ γ := -⟨λa, g (f a) a, - λ c, f.measurable_set_cut (λ a b, g b a = c) $ λ b, (g b).measurable_set_preimage {c}, - (f.finite_range.bUnion (λ b _, (g b).finite_range)).subset $ - by rintro _ ⟨a, rfl⟩; simp; exact ⟨a, a, rfl⟩⟩ - -@[simp] theorem bind_apply (f : α →ₛ β) (g : β → α →ₛ γ) (a) : - f.bind g a = g (f a) a := rfl - -/-- Given a function `g : β → γ` and a simple function `f : α →ₛ β`, `f.map g` return the simple - function `g ∘ f : α →ₛ γ` -/ -def map (g : β → γ) (f : α →ₛ β) : α →ₛ γ := bind f (const α ∘ g) - -theorem map_apply (g : β → γ) (f : α →ₛ β) (a) : f.map g a = g (f a) := rfl - -theorem map_map (g : β → γ) (h: γ → δ) (f : α →ₛ β) : (f.map g).map h = f.map (h ∘ g) := rfl - -@[simp] theorem coe_map (g : β → γ) (f : α →ₛ β) : (f.map g : α → γ) = g ∘ f := rfl - -@[simp] theorem range_map [decidable_eq γ] (g : β → γ) (f : α →ₛ β) : - (f.map g).range = f.range.image g := -finset.coe_injective $ by simp only [coe_range, coe_map, finset.coe_image, range_comp] - -@[simp] theorem map_const (g : β → γ) (b : β) : (const α b).map g = const α (g b) := rfl - -lemma map_preimage (f : α →ₛ β) (g : β → γ) (s : set γ) : - (f.map g) ⁻¹' s = f ⁻¹' ↑(f.range.filter (λb, g b ∈ s)) := -by { simp only [coe_range, sep_mem_eq, set.mem_range, function.comp_app, coe_map, finset.coe_filter, - ← mem_preimage, inter_comm, preimage_inter_range], apply preimage_comp } - -lemma map_preimage_singleton (f : α →ₛ β) (g : β → γ) (c : γ) : - (f.map g) ⁻¹' {c} = f ⁻¹' ↑(f.range.filter (λ b, g b = c)) := -map_preimage _ _ _ - -/-- Composition of a `simple_fun` and a measurable function is a `simple_func`. -/ -def comp [measurable_space β] (f : β →ₛ γ) (g : α → β) (hgm : measurable g) : α →ₛ γ := -{ to_fun := f ∘ g, - finite_range' := f.finite_range.subset $ set.range_comp_subset_range _ _, - measurable_set_fiber' := λ z, hgm (f.measurable_set_fiber z) } - -@[simp] lemma coe_comp [measurable_space β] (f : β →ₛ γ) {g : α → β} (hgm : measurable g) : - ⇑(f.comp g hgm) = f ∘ g := -rfl - -lemma range_comp_subset_range [measurable_space β] (f : β →ₛ γ) {g : α → β} (hgm : measurable g) : - (f.comp g hgm).range ⊆ f.range := -finset.coe_subset.1 $ by simp only [coe_range, coe_comp, set.range_comp_subset_range] - -/-- Extend a `simple_func` along a measurable embedding: `f₁.extend g hg f₂` is the function -`F : β →ₛ γ` such that `F ∘ g = f₁` and `F y = f₂ y` whenever `y ∉ range g`. -/ -def extend [measurable_space β] (f₁ : α →ₛ γ) (g : α → β) - (hg : measurable_embedding g) (f₂ : β →ₛ γ) : β →ₛ γ := -{ to_fun := function.extend g f₁ f₂, - finite_range' := (f₁.finite_range.union $ f₂.finite_range.subset - (image_subset_range _ _)).subset (range_extend_subset _ _ _), - measurable_set_fiber' := - begin - letI : measurable_space γ := ⊤, haveI : measurable_singleton_class γ := ⟨λ _, trivial⟩, - exact λ x, hg.measurable_extend f₁.measurable f₂.measurable (measurable_set_singleton _) - end } - -@[simp] lemma extend_apply [measurable_space β] (f₁ : α →ₛ γ) {g : α → β} - (hg : measurable_embedding g) (f₂ : β →ₛ γ) (x : α) : (f₁.extend g hg f₂) (g x) = f₁ x := -hg.injective.extend_apply _ _ _ - -@[simp] lemma extend_apply' [measurable_space β] (f₁ : α →ₛ γ) {g : α → β} - (hg : measurable_embedding g) (f₂ : β →ₛ γ) {y : β} (h : ¬∃ x, g x = y) : - (f₁.extend g hg f₂) y = f₂ y := -function.extend_apply' _ _ _ h - -@[simp] lemma extend_comp_eq' [measurable_space β] (f₁ : α →ₛ γ) {g : α → β} - (hg : measurable_embedding g) (f₂ : β →ₛ γ) : (f₁.extend g hg f₂) ∘ g = f₁ := -funext $ λ x, extend_apply _ _ _ _ - -@[simp] lemma extend_comp_eq [measurable_space β] (f₁ : α →ₛ γ) {g : α → β} - (hg : measurable_embedding g) (f₂ : β →ₛ γ) : (f₁.extend g hg f₂).comp g hg.measurable = f₁ := -coe_injective $ extend_comp_eq' _ _ _ - -/-- If `f` is a simple function taking values in `β → γ` and `g` is another simple function -with the same domain and codomain `β`, then `f.seq g = f a (g a)`. -/ -def seq (f : α →ₛ (β → γ)) (g : α →ₛ β) : α →ₛ γ := f.bind (λf, g.map f) - -@[simp] lemma seq_apply (f : α →ₛ (β → γ)) (g : α →ₛ β) (a : α) : f.seq g a = f a (g a) := rfl - -/-- Combine two simple functions `f : α →ₛ β` and `g : α →ₛ β` -into `λ a, (f a, g a)`. -/ -def pair (f : α →ₛ β) (g : α →ₛ γ) : α →ₛ (β × γ) := (f.map prod.mk).seq g - -@[simp] lemma pair_apply (f : α →ₛ β) (g : α →ₛ γ) (a) : pair f g a = (f a, g a) := rfl - -lemma pair_preimage (f : α →ₛ β) (g : α →ₛ γ) (s : set β) (t : set γ) : - pair f g ⁻¹' s ×ˢ t = (f ⁻¹' s) ∩ (g ⁻¹' t) := rfl - -/- A special form of `pair_preimage` -/ -lemma pair_preimage_singleton (f : α →ₛ β) (g : α →ₛ γ) (b : β) (c : γ) : - (pair f g) ⁻¹' {(b, c)} = (f ⁻¹' {b}) ∩ (g ⁻¹' {c}) := -by { rw ← singleton_prod_singleton, exact pair_preimage _ _ _ _ } - -theorem bind_const (f : α →ₛ β) : f.bind (const α) = f := by ext; simp - -@[to_additive] instance [has_one β] : has_one (α →ₛ β) := ⟨const α 1⟩ -@[to_additive] instance [has_mul β] : has_mul (α →ₛ β) := ⟨λf g, (f.map (*)).seq g⟩ -@[to_additive] instance [has_div β] : has_div (α →ₛ β) := ⟨λf g, (f.map (/)).seq g⟩ -@[to_additive] instance [has_inv β] : has_inv (α →ₛ β) := ⟨λf, f.map (has_inv.inv)⟩ -instance [has_sup β] : has_sup (α →ₛ β) := ⟨λf g, (f.map (⊔)).seq g⟩ -instance [has_inf β] : has_inf (α →ₛ β) := ⟨λf g, (f.map (⊓)).seq g⟩ -instance [has_le β] : has_le (α →ₛ β) := ⟨λf g, ∀a, f a ≤ g a⟩ - -@[simp, to_additive] lemma const_one [has_one β] : const α (1 : β) = 1 := rfl - -@[simp, norm_cast, to_additive] lemma coe_one [has_one β] : ⇑(1 : α →ₛ β) = 1 := rfl -@[simp, norm_cast, to_additive] lemma coe_mul [has_mul β] (f g : α →ₛ β) : ⇑(f * g) = f * g := rfl -@[simp, norm_cast, to_additive] lemma coe_inv [has_inv β] (f : α →ₛ β) : ⇑(f⁻¹) = f⁻¹ := rfl -@[simp, norm_cast, to_additive] lemma coe_div [has_div β] (f g : α →ₛ β) : ⇑(f / g) = f / g := rfl -@[simp, norm_cast] lemma coe_le [preorder β] {f g : α →ₛ β} : (f : α → β) ≤ g ↔ f ≤ g := iff.rfl -@[simp, norm_cast] lemma coe_sup [has_sup β] (f g : α →ₛ β) : ⇑(f ⊔ g) = f ⊔ g := rfl -@[simp, norm_cast] lemma coe_inf [has_inf β] (f g : α →ₛ β) : ⇑(f ⊓ g) = f ⊓ g := rfl - -@[to_additive] lemma mul_apply [has_mul β] (f g : α →ₛ β) (a : α) : (f * g) a = f a * g a := rfl -@[to_additive] lemma div_apply [has_div β] (f g : α →ₛ β) (x : α) : (f / g) x = f x / g x := rfl -@[to_additive] lemma inv_apply [has_inv β] (f : α →ₛ β) (x : α) : f⁻¹ x = (f x)⁻¹ := rfl -lemma sup_apply [has_sup β] (f g : α →ₛ β) (a : α) : (f ⊔ g) a = f a ⊔ g a := rfl -lemma inf_apply [has_inf β] (f g : α →ₛ β) (a : α) : (f ⊓ g) a = f a ⊓ g a := rfl - -@[simp, to_additive] lemma range_one [nonempty α] [has_one β] : (1 : α →ₛ β).range = {1} := -finset.ext $ λ x, by simp [eq_comm] - -@[simp] lemma range_eq_empty_of_is_empty {β} [hα : is_empty α] (f : α →ₛ β) : - f.range = ∅ := -begin - rw ← finset.not_nonempty_iff_eq_empty, - by_contra, - obtain ⟨y, hy_mem⟩ := h, - rw [simple_func.mem_range, set.mem_range] at hy_mem, - obtain ⟨x, hxy⟩ := hy_mem, - rw is_empty_iff at hα, - exact hα x, -end - -lemma eq_zero_of_mem_range_zero [has_zero β] : ∀ {y : β}, y ∈ (0 : α →ₛ β).range → y = 0 := -forall_range_iff.2 $ λ x, rfl - -@[to_additive] -lemma mul_eq_map₂ [has_mul β] (f g : α →ₛ β) : f * g = (pair f g).map (λp:β×β, p.1 * p.2) := rfl - -lemma sup_eq_map₂ [has_sup β] (f g : α →ₛ β) : f ⊔ g = (pair f g).map (λp:β×β, p.1 ⊔ p.2) := rfl - -@[to_additive] -lemma const_mul_eq_map [has_mul β] (f : α →ₛ β) (b : β) : const α b * f = f.map (λa, b * a) := rfl - -@[to_additive] -theorem map_mul [has_mul β] [has_mul γ] {g : β → γ} - (hg : ∀ x y, g (x * y) = g x * g y) (f₁ f₂ : α →ₛ β) : (f₁ * f₂).map g = f₁.map g * f₂.map g := -ext $ λ x, hg _ _ - -variables {K : Type*} - -instance [has_smul K β] : has_smul K (α →ₛ β) := ⟨λ k f, f.map ((•) k)⟩ -@[simp] lemma coe_smul [has_smul K β] (c : K) (f : α →ₛ β) : ⇑(c • f) = c • f := rfl -lemma smul_apply [has_smul K β] (k : K) (f : α →ₛ β) (a : α) : (k • f) a = k • f a := rfl - -instance has_nat_pow [monoid β] : has_pow (α →ₛ β) ℕ := ⟨λ f n, f.map (^ n)⟩ -@[simp] lemma coe_pow [monoid β] (f : α →ₛ β) (n : ℕ) : ⇑(f ^ n) = f ^ n := rfl -lemma pow_apply [monoid β] (n : ℕ) (f : α →ₛ β) (a : α) : (f ^ n) a = f a ^ n := rfl - -instance has_int_pow [div_inv_monoid β] : has_pow (α →ₛ β) ℤ := ⟨λ f n, f.map (^ n)⟩ -@[simp] lemma coe_zpow [div_inv_monoid β] (f : α →ₛ β) (z : ℤ) : ⇑(f ^ z) = f ^ z := rfl -lemma zpow_apply [div_inv_monoid β] (z : ℤ) (f : α →ₛ β) (a : α) : (f ^ z) a = f a ^ z := rfl - --- TODO: work out how to generate these instances with `to_additive`, which gets confused by the --- argument order swap between `coe_smul` and `coe_pow`. -section additive - -instance [add_monoid β] : add_monoid (α →ₛ β) := -function.injective.add_monoid (λ f, show α → β, from f) coe_injective coe_zero coe_add - (λ _ _, coe_smul _ _) - -instance [add_comm_monoid β] : add_comm_monoid (α →ₛ β) := -function.injective.add_comm_monoid (λ f, show α → β, from f) coe_injective coe_zero coe_add - (λ _ _, coe_smul _ _) - -instance [add_group β] : add_group (α →ₛ β) := -function.injective.add_group (λ f, show α → β, from f) coe_injective - coe_zero coe_add coe_neg coe_sub (λ _ _, coe_smul _ _) (λ _ _, coe_smul _ _) - -instance [add_comm_group β] : add_comm_group (α →ₛ β) := -function.injective.add_comm_group (λ f, show α → β, from f) coe_injective - coe_zero coe_add coe_neg coe_sub (λ _ _, coe_smul _ _) (λ _ _, coe_smul _ _) - -end additive - -@[to_additive] instance [monoid β] : monoid (α →ₛ β) := -function.injective.monoid (λ f, show α → β, from f) coe_injective coe_one coe_mul coe_pow - -@[to_additive] instance [comm_monoid β] : comm_monoid (α →ₛ β) := -function.injective.comm_monoid (λ f, show α → β, from f) coe_injective coe_one coe_mul coe_pow - -@[to_additive] instance [group β] : group (α →ₛ β) := -function.injective.group (λ f, show α → β, from f) coe_injective - coe_one coe_mul coe_inv coe_div coe_pow coe_zpow - -@[to_additive] instance [comm_group β] : comm_group (α →ₛ β) := -function.injective.comm_group (λ f, show α → β, from f) coe_injective - coe_one coe_mul coe_inv coe_div coe_pow coe_zpow - -instance [semiring K] [add_comm_monoid β] [module K β] : module K (α →ₛ β) := -function.injective.module K ⟨λ f, show α → β, from f, coe_zero, coe_add⟩ - coe_injective coe_smul - -lemma smul_eq_map [has_smul K β] (k : K) (f : α →ₛ β) : k • f = f.map ((•) k) := rfl - -instance [preorder β] : preorder (α →ₛ β) := -{ le_refl := λf a, le_rfl, - le_trans := λf g h hfg hgh a, le_trans (hfg _) (hgh a), - .. simple_func.has_le } - -instance [partial_order β] : partial_order (α →ₛ β) := -{ le_antisymm := assume f g hfg hgf, ext $ assume a, le_antisymm (hfg a) (hgf a), - .. simple_func.preorder } - -instance [has_le β] [order_bot β] : order_bot (α →ₛ β) := -{ bot := const α ⊥, bot_le := λf a, bot_le } - -instance [has_le β] [order_top β] : order_top (α →ₛ β) := -{ top := const α ⊤, le_top := λf a, le_top } - -instance [semilattice_inf β] : semilattice_inf (α →ₛ β) := -{ inf := (⊓), - inf_le_left := assume f g a, inf_le_left, - inf_le_right := assume f g a, inf_le_right, - le_inf := assume f g h hfh hgh a, le_inf (hfh a) (hgh a), - .. simple_func.partial_order } - -instance [semilattice_sup β] : semilattice_sup (α →ₛ β) := -{ sup := (⊔), - le_sup_left := assume f g a, le_sup_left, - le_sup_right := assume f g a, le_sup_right, - sup_le := assume f g h hfh hgh a, sup_le (hfh a) (hgh a), - .. simple_func.partial_order } - -instance [lattice β] : lattice (α →ₛ β) := -{ .. simple_func.semilattice_sup,.. simple_func.semilattice_inf } - -instance [has_le β] [bounded_order β] : bounded_order (α →ₛ β) := -{ .. simple_func.order_bot, .. simple_func.order_top } - -lemma finset_sup_apply [semilattice_sup β] [order_bot β] {f : γ → α →ₛ β} (s : finset γ) (a : α) : - s.sup f a = s.sup (λc, f c a) := -begin - refine finset.induction_on s rfl _, - assume a s hs ih, - rw [finset.sup_insert, finset.sup_insert, sup_apply, ih] -end - -section restrict - -variables [has_zero β] - -/-- Restrict a simple function `f : α →ₛ β` to a set `s`. If `s` is measurable, -then `f.restrict s a = if a ∈ s then f a else 0`, otherwise `f.restrict s = const α 0`. -/ -def restrict (f : α →ₛ β) (s : set α) : α →ₛ β := -if hs : measurable_set s then piecewise s hs f 0 else 0 - -theorem restrict_of_not_measurable {f : α →ₛ β} {s : set α} - (hs : ¬measurable_set s) : - restrict f s = 0 := -dif_neg hs - -@[simp] theorem coe_restrict (f : α →ₛ β) {s : set α} (hs : measurable_set s) : - ⇑(restrict f s) = indicator s f := -by { rw [restrict, dif_pos hs], refl } - -@[simp] theorem restrict_univ (f : α →ₛ β) : restrict f univ = f := -by simp [restrict] - -@[simp] theorem restrict_empty (f : α →ₛ β) : restrict f ∅ = 0 := -by simp [restrict] - -theorem map_restrict_of_zero [has_zero γ] {g : β → γ} (hg : g 0 = 0) (f : α →ₛ β) (s : set α) : - (f.restrict s).map g = (f.map g).restrict s := -ext $ λ x, -if hs : measurable_set s then by simp [hs, set.indicator_comp_of_zero hg] -else by simp [restrict_of_not_measurable hs, hg] - -theorem map_coe_ennreal_restrict (f : α →ₛ ℝ≥0) (s : set α) : - (f.restrict s).map (coe : ℝ≥0 → ℝ≥0∞) = (f.map coe).restrict s := -map_restrict_of_zero ennreal.coe_zero _ _ - -theorem map_coe_nnreal_restrict (f : α →ₛ ℝ≥0) (s : set α) : - (f.restrict s).map (coe : ℝ≥0 → ℝ) = (f.map coe).restrict s := -map_restrict_of_zero nnreal.coe_zero _ _ - -theorem restrict_apply (f : α →ₛ β) {s : set α} (hs : measurable_set s) (a) : - restrict f s a = indicator s f a := -by simp only [f.coe_restrict hs] - -theorem restrict_preimage (f : α →ₛ β) {s : set α} (hs : measurable_set s) - {t : set β} (ht : (0:β) ∉ t) : restrict f s ⁻¹' t = s ∩ f ⁻¹' t := -by simp [hs, indicator_preimage_of_not_mem _ _ ht, inter_comm] - -theorem restrict_preimage_singleton (f : α →ₛ β) {s : set α} (hs : measurable_set s) - {r : β} (hr : r ≠ 0) : restrict f s ⁻¹' {r} = s ∩ f ⁻¹' {r} := -f.restrict_preimage hs hr.symm - -lemma mem_restrict_range {r : β} {s : set α} {f : α →ₛ β} (hs : measurable_set s) : - r ∈ (restrict f s).range ↔ (r = 0 ∧ s ≠ univ) ∨ (r ∈ f '' s) := -by rw [← finset.mem_coe, coe_range, coe_restrict _ hs, mem_range_indicator] - -lemma mem_image_of_mem_range_restrict {r : β} {s : set α} {f : α →ₛ β} - (hr : r ∈ (restrict f s).range) (h0 : r ≠ 0) : - r ∈ f '' s := -if hs : measurable_set s then by simpa [mem_restrict_range hs, h0] using hr -else by { rw [restrict_of_not_measurable hs] at hr, - exact (h0 $ eq_zero_of_mem_range_zero hr).elim } - -@[mono] lemma restrict_mono [preorder β] (s : set α) {f g : α →ₛ β} (H : f ≤ g) : - f.restrict s ≤ g.restrict s := -if hs : measurable_set s then λ x, by simp only [coe_restrict _ hs, indicator_le_indicator (H x)] -else by simp only [restrict_of_not_measurable hs, le_refl] - -end restrict - -section approx - -section -variables [semilattice_sup β] [order_bot β] [has_zero β] - -/-- Fix a sequence `i : ℕ → β`. Given a function `α → β`, its `n`-th approximation -by simple functions is defined so that in case `β = ℝ≥0∞` it sends each `a` to the supremum -of the set `{i k | k ≤ n ∧ i k ≤ f a}`, see `approx_apply` and `supr_approx_apply` for details. -/ -def approx (i : ℕ → β) (f : α → β) (n : ℕ) : α →ₛ β := -(finset.range n).sup (λk, restrict (const α (i k)) {a:α | i k ≤ f a}) - -lemma approx_apply [topological_space β] [order_closed_topology β] [measurable_space β] - [opens_measurable_space β] {i : ℕ → β} {f : α → β} {n : ℕ} (a : α) (hf : measurable f) : - (approx i f n : α →ₛ β) a = (finset.range n).sup (λk, if i k ≤ f a then i k else 0) := -begin - dsimp only [approx], - rw [finset_sup_apply], - congr, - funext k, - rw [restrict_apply], - refl, - exact (hf measurable_set_Ici) -end - -lemma monotone_approx (i : ℕ → β) (f : α → β) : monotone (approx i f) := -assume n m h, finset.sup_mono $ finset.range_subset.2 h - -lemma approx_comp [topological_space β] [order_closed_topology β] [measurable_space β] - [opens_measurable_space β] [measurable_space γ] - {i : ℕ → β} {f : γ → β} {g : α → γ} {n : ℕ} (a : α) - (hf : measurable f) (hg : measurable g) : - (approx i (f ∘ g) n : α →ₛ β) a = (approx i f n : γ →ₛ β) (g a) := -by rw [approx_apply _ hf, approx_apply _ (hf.comp hg)] - -end - -lemma supr_approx_apply [topological_space β] [complete_lattice β] [order_closed_topology β] - [has_zero β] [measurable_space β] [opens_measurable_space β] - (i : ℕ → β) (f : α → β) (a : α) (hf : measurable f) (h_zero : (0 : β) = ⊥) : - (⨆n, (approx i f n : α →ₛ β) a) = (⨆k (h : i k ≤ f a), i k) := -begin - refine le_antisymm (supr_le $ assume n, _) (supr_le $ assume k, supr_le $ assume hk, _), - { rw [approx_apply a hf, h_zero], - refine finset.sup_le (assume k hk, _), - split_ifs, - exact le_supr_of_le k (le_supr _ h), - exact bot_le }, - { refine le_supr_of_le (k+1) _, - rw [approx_apply a hf], - have : k ∈ finset.range (k+1) := finset.mem_range.2 (nat.lt_succ_self _), - refine le_trans (le_of_eq _) (finset.le_sup this), - rw [if_pos hk] } -end - -end approx - -section eapprox - -/-- A sequence of `ℝ≥0∞`s such that its range is the set of non-negative rational numbers. -/ -def ennreal_rat_embed (n : ℕ) : ℝ≥0∞ := -ennreal.of_real ((encodable.decode ℚ n).get_or_else (0 : ℚ)) - -lemma ennreal_rat_embed_encode (q : ℚ) : - ennreal_rat_embed (encodable.encode q) = real.to_nnreal q := -by rw [ennreal_rat_embed, encodable.encodek]; refl - -/-- Approximate a function `α → ℝ≥0∞` by a sequence of simple functions. -/ -def eapprox : (α → ℝ≥0∞) → ℕ → α →ₛ ℝ≥0∞ := -approx ennreal_rat_embed - -lemma eapprox_lt_top (f : α → ℝ≥0∞) (n : ℕ) (a : α) : eapprox f n a < ∞ := -begin - simp only [eapprox, approx, finset_sup_apply, finset.sup_lt_iff, with_top.zero_lt_top, - finset.mem_range, ennreal.bot_eq_zero, restrict], - assume b hb, - split_ifs, - { simp only [coe_zero, coe_piecewise, piecewise_eq_indicator, coe_const], - calc {a : α | ennreal_rat_embed b ≤ f a}.indicator (λ x, ennreal_rat_embed b) a - ≤ ennreal_rat_embed b : indicator_le_self _ _ a - ... < ⊤ : ennreal.coe_lt_top }, - { exact with_top.zero_lt_top }, -end - -@[mono] lemma monotone_eapprox (f : α → ℝ≥0∞) : monotone (eapprox f) := -monotone_approx _ f - -lemma supr_eapprox_apply (f : α → ℝ≥0∞) (hf : measurable f) (a : α) : - (⨆n, (eapprox f n : α →ₛ ℝ≥0∞) a) = f a := -begin - rw [eapprox, supr_approx_apply ennreal_rat_embed f a hf rfl], - refine le_antisymm (supr_le $ assume i, supr_le $ assume hi, hi) (le_of_not_gt _), - assume h, - rcases ennreal.lt_iff_exists_rat_btwn.1 h with ⟨q, hq, lt_q, q_lt⟩, - have : (real.to_nnreal q : ℝ≥0∞) ≤ - (⨆ (k : ℕ) (h : ennreal_rat_embed k ≤ f a), ennreal_rat_embed k), - { refine le_supr_of_le (encodable.encode q) _, - rw [ennreal_rat_embed_encode q], - refine le_supr_of_le (le_of_lt q_lt) _, - exact le_rfl }, - exact lt_irrefl _ (lt_of_le_of_lt this lt_q) -end - -lemma eapprox_comp [measurable_space γ] {f : γ → ℝ≥0∞} {g : α → γ} {n : ℕ} - (hf : measurable f) (hg : measurable g) : - (eapprox (f ∘ g) n : α → ℝ≥0∞) = (eapprox f n : γ →ₛ ℝ≥0∞) ∘ g := -funext $ assume a, approx_comp a hf hg - -/-- Approximate a function `α → ℝ≥0∞` by a series of simple functions taking their values -in `ℝ≥0`. -/ -def eapprox_diff (f : α → ℝ≥0∞) : ∀ (n : ℕ), α →ₛ ℝ≥0 -| 0 := (eapprox f 0).map ennreal.to_nnreal -| (n+1) := (eapprox f (n+1) - eapprox f n).map ennreal.to_nnreal - -lemma sum_eapprox_diff (f : α → ℝ≥0∞) (n : ℕ) (a : α) : - (∑ k in finset.range (n+1), (eapprox_diff f k a : ℝ≥0∞)) = eapprox f n a := -begin - induction n with n IH, - { simp only [nat.nat_zero_eq_zero, finset.sum_singleton, finset.range_one], refl }, - { rw [finset.sum_range_succ, nat.succ_eq_add_one, IH, eapprox_diff, coe_map, function.comp_app, - coe_sub, pi.sub_apply, ennreal.coe_to_nnreal, - add_tsub_cancel_of_le (monotone_eapprox f (nat.le_succ _) _)], - apply (lt_of_le_of_lt _ (eapprox_lt_top f (n+1) a)).ne, - rw tsub_le_iff_right, - exact le_self_add }, -end - -lemma tsum_eapprox_diff (f : α → ℝ≥0∞) (hf : measurable f) (a : α) : - (∑' n, (eapprox_diff f n a : ℝ≥0∞)) = f a := -by simp_rw [ennreal.tsum_eq_supr_nat' (tendsto_add_at_top_nat 1), sum_eapprox_diff, - supr_eapprox_apply f hf a] - -end eapprox - -end measurable - -section measure -variables {m : measurable_space α} {μ ν : measure α} - -/-- Integral of a simple function whose codomain is `ℝ≥0∞`. -/ -def lintegral {m : measurable_space α} (f : α →ₛ ℝ≥0∞) (μ : measure α) : ℝ≥0∞ := -∑ x in f.range, x * μ (f ⁻¹' {x}) - -lemma lintegral_eq_of_subset (f : α →ₛ ℝ≥0∞) {s : finset ℝ≥0∞} - (hs : ∀ x, f x ≠ 0 → μ (f ⁻¹' {f x}) ≠ 0 → f x ∈ s) : - f.lintegral μ = ∑ x in s, x * μ (f ⁻¹' {x}) := -begin - refine finset.sum_bij_ne_zero (λr _ _, r) _ _ _ _, - { simpa only [forall_range_iff, mul_ne_zero_iff, and_imp] }, - { intros, assumption }, - { intros b _ hb, - refine ⟨b, _, hb, rfl⟩, - rw [mem_range, ← preimage_singleton_nonempty], - exact nonempty_of_measure_ne_zero (mul_ne_zero_iff.1 hb).2 }, - { intros, refl } -end - -lemma lintegral_eq_of_subset' (f : α →ₛ ℝ≥0∞) {s : finset ℝ≥0∞} - (hs : f.range \ {0} ⊆ s) : - f.lintegral μ = ∑ x in s, x * μ (f ⁻¹' {x}) := -f.lintegral_eq_of_subset $ λ x hfx _, hs $ - finset.mem_sdiff.2 ⟨f.mem_range_self x, mt finset.mem_singleton.1 hfx⟩ - -/-- Calculate the integral of `(g ∘ f)`, where `g : β → ℝ≥0∞` and `f : α →ₛ β`. -/ -lemma map_lintegral (g : β → ℝ≥0∞) (f : α →ₛ β) : - (f.map g).lintegral μ = ∑ x in f.range, g x * μ (f ⁻¹' {x}) := -begin - simp only [lintegral, range_map], - refine finset.sum_image' _ (assume b hb, _), - rcases mem_range.1 hb with ⟨a, rfl⟩, - rw [map_preimage_singleton, ← f.sum_measure_preimage_singleton, finset.mul_sum], - refine finset.sum_congr _ _, - { congr }, - { assume x, simp only [finset.mem_filter], rintro ⟨_, h⟩, rw h }, -end - -lemma add_lintegral (f g : α →ₛ ℝ≥0∞) : (f + g).lintegral μ = f.lintegral μ + g.lintegral μ := -calc (f + g).lintegral μ = - ∑ x in (pair f g).range, (x.1 * μ (pair f g ⁻¹' {x}) + x.2 * μ (pair f g ⁻¹' {x})) : - by rw [add_eq_map₂, map_lintegral]; exact finset.sum_congr rfl (assume a ha, add_mul _ _ _) - ... = ∑ x in (pair f g).range, x.1 * μ (pair f g ⁻¹' {x}) + - ∑ x in (pair f g).range, x.2 * μ (pair f g ⁻¹' {x}) : by rw [finset.sum_add_distrib] - ... = ((pair f g).map prod.fst).lintegral μ + ((pair f g).map prod.snd).lintegral μ : - by rw [map_lintegral, map_lintegral] - ... = lintegral f μ + lintegral g μ : rfl - -lemma const_mul_lintegral (f : α →ₛ ℝ≥0∞) (x : ℝ≥0∞) : - (const α x * f).lintegral μ = x * f.lintegral μ := -calc (f.map (λa, x * a)).lintegral μ = ∑ r in f.range, x * r * μ (f ⁻¹' {r}) : - map_lintegral _ _ - ... = ∑ r in f.range, x * (r * μ (f ⁻¹' {r})) : - finset.sum_congr rfl (assume a ha, mul_assoc _ _ _) - ... = x * f.lintegral μ : - finset.mul_sum.symm - -/-- Integral of a simple function `α →ₛ ℝ≥0∞` as a bilinear map. -/ -def lintegralₗ {m : measurable_space α} : (α →ₛ ℝ≥0∞) →ₗ[ℝ≥0∞] measure α →ₗ[ℝ≥0∞] ℝ≥0∞ := -{ to_fun := λ f, - { to_fun := lintegral f, - map_add' := by simp [lintegral, mul_add, finset.sum_add_distrib], - map_smul' := λ c μ, by simp [lintegral, mul_left_comm _ c, finset.mul_sum] }, - map_add' := λ f g, linear_map.ext (λ μ, add_lintegral f g), - map_smul' := λ c f, linear_map.ext (λ μ, const_mul_lintegral f c) } - -@[simp] lemma zero_lintegral : (0 : α →ₛ ℝ≥0∞).lintegral μ = 0 := -linear_map.ext_iff.1 lintegralₗ.map_zero μ - -lemma lintegral_add {ν} (f : α →ₛ ℝ≥0∞) : f.lintegral (μ + ν) = f.lintegral μ + f.lintegral ν := -(lintegralₗ f).map_add μ ν - -lemma lintegral_smul (f : α →ₛ ℝ≥0∞) (c : ℝ≥0∞) : - f.lintegral (c • μ) = c • f.lintegral μ := -(lintegralₗ f).map_smul c μ - -@[simp] lemma lintegral_zero [measurable_space α] (f : α →ₛ ℝ≥0∞) : - f.lintegral 0 = 0 := -(lintegralₗ f).map_zero - -lemma lintegral_sum {m : measurable_space α} {ι} (f : α →ₛ ℝ≥0∞) (μ : ι → measure α) : - f.lintegral (measure.sum μ) = ∑' i, f.lintegral (μ i) := -begin - simp only [lintegral, measure.sum_apply, f.measurable_set_preimage, ← finset.tsum_subtype, - ← ennreal.tsum_mul_left], - apply ennreal.tsum_comm -end - -lemma restrict_lintegral (f : α →ₛ ℝ≥0∞) {s : set α} (hs : measurable_set s) : - (restrict f s).lintegral μ = ∑ r in f.range, r * μ (f ⁻¹' {r} ∩ s) := -calc (restrict f s).lintegral μ = ∑ r in f.range, r * μ (restrict f s ⁻¹' {r}) : - lintegral_eq_of_subset _ $ λ x hx, if hxs : x ∈ s - then λ _, by simp only [f.restrict_apply hs, indicator_of_mem hxs, mem_range_self] - else false.elim $ hx $ by simp [*] -... = ∑ r in f.range, r * μ (f ⁻¹' {r} ∩ s) : - finset.sum_congr rfl $ forall_range_iff.2 $ λ b, if hb : f b = 0 then by simp only [hb, zero_mul] - else by rw [restrict_preimage_singleton _ hs hb, inter_comm] - -lemma lintegral_restrict {m : measurable_space α} (f : α →ₛ ℝ≥0∞) (s : set α) (μ : measure α) : - f.lintegral (μ.restrict s) = ∑ y in f.range, y * μ (f ⁻¹' {y} ∩ s) := -by simp only [lintegral, measure.restrict_apply, f.measurable_set_preimage] - -lemma restrict_lintegral_eq_lintegral_restrict (f : α →ₛ ℝ≥0∞) {s : set α} - (hs : measurable_set s) : - (restrict f s).lintegral μ = f.lintegral (μ.restrict s) := -by rw [f.restrict_lintegral hs, lintegral_restrict] - -lemma const_lintegral (c : ℝ≥0∞) : (const α c).lintegral μ = c * μ univ := -begin - rw [lintegral], - casesI is_empty_or_nonempty α, - { simp [μ.eq_zero_of_is_empty] }, - { simp [preimage_const_of_mem] }, -end - -lemma const_lintegral_restrict (c : ℝ≥0∞) (s : set α) : - (const α c).lintegral (μ.restrict s) = c * μ s := -by rw [const_lintegral, measure.restrict_apply measurable_set.univ, univ_inter] - -lemma restrict_const_lintegral (c : ℝ≥0∞) {s : set α} (hs : measurable_set s) : - ((const α c).restrict s).lintegral μ = c * μ s := -by rw [restrict_lintegral_eq_lintegral_restrict _ hs, const_lintegral_restrict] - -lemma le_sup_lintegral (f g : α →ₛ ℝ≥0∞) : f.lintegral μ ⊔ g.lintegral μ ≤ (f ⊔ g).lintegral μ := -calc f.lintegral μ ⊔ g.lintegral μ = - ((pair f g).map prod.fst).lintegral μ ⊔ ((pair f g).map prod.snd).lintegral μ : rfl - ... ≤ ∑ x in (pair f g).range, (x.1 ⊔ x.2) * μ (pair f g ⁻¹' {x}) : - begin - rw [map_lintegral, map_lintegral], - refine sup_le _ _; - refine finset.sum_le_sum (λ a _, mul_le_mul_right' _ _), - exact le_sup_left, - exact le_sup_right - end - ... = (f ⊔ g).lintegral μ : by rw [sup_eq_map₂, map_lintegral] - -/-- `simple_func.lintegral` is monotone both in function and in measure. -/ -@[mono] lemma lintegral_mono {f g : α →ₛ ℝ≥0∞} (hfg : f ≤ g) (hμν : μ ≤ ν) : - f.lintegral μ ≤ g.lintegral ν := -calc f.lintegral μ ≤ f.lintegral μ ⊔ g.lintegral μ : le_sup_left - ... ≤ (f ⊔ g).lintegral μ : le_sup_lintegral _ _ - ... = g.lintegral μ : by rw [sup_of_le_right hfg] - ... ≤ g.lintegral ν : finset.sum_le_sum $ λ y hy, ennreal.mul_left_mono $ - hμν _ (g.measurable_set_preimage _) - -/-- `simple_func.lintegral` depends only on the measures of `f ⁻¹' {y}`. -/ -lemma lintegral_eq_of_measure_preimage [measurable_space β] {f : α →ₛ ℝ≥0∞} {g : β →ₛ ℝ≥0∞} - {ν : measure β} (H : ∀ y, μ (f ⁻¹' {y}) = ν (g ⁻¹' {y})) : - f.lintegral μ = g.lintegral ν := -begin - simp only [lintegral, ← H], - apply lintegral_eq_of_subset, - simp only [H], - intros, - exact mem_range_of_measure_ne_zero ‹_› -end - -/-- If two simple functions are equal a.e., then their `lintegral`s are equal. -/ -lemma lintegral_congr {f g : α →ₛ ℝ≥0∞} (h : f =ᵐ[μ] g) : - f.lintegral μ = g.lintegral μ := -lintegral_eq_of_measure_preimage $ λ y, measure_congr $ - eventually.set_eq $ h.mono $ λ x hx, by simp [hx] - -lemma lintegral_map' {β} [measurable_space β] {μ' : measure β} (f : α →ₛ ℝ≥0∞) (g : β →ₛ ℝ≥0∞) - (m' : α → β) (eq : ∀ a, f a = g (m' a)) (h : ∀s, measurable_set s → μ' s = μ (m' ⁻¹' s)) : - f.lintegral μ = g.lintegral μ' := -lintegral_eq_of_measure_preimage $ λ y, -by { simp only [preimage, eq], exact (h (g ⁻¹' {y}) (g.measurable_set_preimage _)).symm } - -lemma lintegral_map {β} [measurable_space β] (g : β →ₛ ℝ≥0∞) {f : α → β} (hf : measurable f) : - g.lintegral (measure.map f μ) = (g.comp f hf).lintegral μ := -eq.symm $ lintegral_map' _ _ f (λ a, rfl) (λ s hs, measure.map_apply hf hs) - -end measure - -section fin_meas_supp - -open finset function - -lemma support_eq [measurable_space α] [has_zero β] (f : α →ₛ β) : - support f = ⋃ y ∈ f.range.filter (λ y, y ≠ 0), f ⁻¹' {y} := -set.ext $ λ x, by simp only [mem_support, set.mem_preimage, mem_filter, mem_range_self, true_and, - exists_prop, mem_Union, set.mem_range, mem_singleton_iff, exists_eq_right'] - -variables {m : measurable_space α} [has_zero β] [has_zero γ] {μ : measure α} {f : α →ₛ β} - -lemma measurable_set_support [measurable_space α] (f : α →ₛ β) : measurable_set (support f) := -by { rw f.support_eq, exact finset.measurable_set_bUnion _ (λ y hy, measurable_set_fiber _ _), } - -/-- A `simple_func` has finite measure support if it is equal to `0` outside of a set of finite -measure. -/ -protected def fin_meas_supp {m : measurable_space α} (f : α →ₛ β) (μ : measure α) : Prop := -f =ᶠ[μ.cofinite] 0 - -lemma fin_meas_supp_iff_support : f.fin_meas_supp μ ↔ μ (support f) < ∞ := iff.rfl - -lemma fin_meas_supp_iff : f.fin_meas_supp μ ↔ ∀ y ≠ 0, μ (f ⁻¹' {y}) < ∞ := -begin - split, - { refine λ h y hy, lt_of_le_of_lt (measure_mono _) h, - exact λ x hx (H : f x = 0), hy $ H ▸ eq.symm hx }, - { intro H, - rw [fin_meas_supp_iff_support, support_eq], - refine lt_of_le_of_lt (measure_bUnion_finset_le _ _) (sum_lt_top _), - exact λ y hy, (H y (finset.mem_filter.1 hy).2).ne } -end - -namespace fin_meas_supp - -lemma meas_preimage_singleton_ne_zero (h : f.fin_meas_supp μ) {y : β} (hy : y ≠ 0) : - μ (f ⁻¹' {y}) < ∞ := -fin_meas_supp_iff.1 h y hy - -protected lemma map {g : β → γ} (hf : f.fin_meas_supp μ) (hg : g 0 = 0) : - (f.map g).fin_meas_supp μ := -flip lt_of_le_of_lt hf (measure_mono $ support_comp_subset hg f) - -lemma of_map {g : β → γ} (h : (f.map g).fin_meas_supp μ) (hg : ∀b, g b = 0 → b = 0) : - f.fin_meas_supp μ := -flip lt_of_le_of_lt h $ measure_mono $ support_subset_comp hg _ - -lemma map_iff {g : β → γ} (hg : ∀ {b}, g b = 0 ↔ b = 0) : - (f.map g).fin_meas_supp μ ↔ f.fin_meas_supp μ := -⟨λ h, h.of_map $ λ b, hg.1, λ h, h.map $ hg.2 rfl⟩ - -protected lemma pair {g : α →ₛ γ} (hf : f.fin_meas_supp μ) (hg : g.fin_meas_supp μ) : - (pair f g).fin_meas_supp μ := -calc μ (support $ pair f g) = μ (support f ∪ support g) : congr_arg μ $ support_prod_mk f g -... ≤ μ (support f) + μ (support g) : measure_union_le _ _ -... < _ : add_lt_top.2 ⟨hf, hg⟩ - -protected lemma map₂ [has_zero δ] (hf : f.fin_meas_supp μ) - {g : α →ₛ γ} (hg : g.fin_meas_supp μ) {op : β → γ → δ} (H : op 0 0 = 0) : - ((pair f g).map (function.uncurry op)).fin_meas_supp μ := -(hf.pair hg).map H - -protected lemma add {β} [add_monoid β] {f g : α →ₛ β} (hf : f.fin_meas_supp μ) - (hg : g.fin_meas_supp μ) : - (f + g).fin_meas_supp μ := -by { rw [add_eq_map₂], exact hf.map₂ hg (zero_add 0) } - -protected lemma mul {β} [monoid_with_zero β] {f g : α →ₛ β} (hf : f.fin_meas_supp μ) - (hg : g.fin_meas_supp μ) : - (f * g).fin_meas_supp μ := -by { rw [mul_eq_map₂], exact hf.map₂ hg (zero_mul 0) } - -lemma lintegral_lt_top {f : α →ₛ ℝ≥0∞} (hm : f.fin_meas_supp μ) (hf : ∀ᵐ a ∂μ, f a ≠ ∞) : - f.lintegral μ < ∞ := -begin - refine sum_lt_top (λ a ha, _), - rcases eq_or_ne a ∞ with rfl|ha, - { simp only [ae_iff, ne.def, not_not] at hf, - simp [set.preimage, hf] }, - { by_cases ha0 : a = 0, - { subst a, rwa [zero_mul] }, - { exact mul_ne_top ha (fin_meas_supp_iff.1 hm _ ha0).ne } } -end - -lemma of_lintegral_ne_top {f : α →ₛ ℝ≥0∞} (h : f.lintegral μ ≠ ∞) : f.fin_meas_supp μ := -begin - refine fin_meas_supp_iff.2 (λ b hb, _), - rw [f.lintegral_eq_of_subset' (finset.subset_insert b _)] at h, - refine ennreal.lt_top_of_mul_ne_top_right _ hb, - exact (lt_top_of_sum_ne_top h (finset.mem_insert_self _ _)).ne -end - -lemma iff_lintegral_lt_top {f : α →ₛ ℝ≥0∞} (hf : ∀ᵐ a ∂μ, f a ≠ ∞) : - f.fin_meas_supp μ ↔ f.lintegral μ < ∞ := -⟨λ h, h.lintegral_lt_top hf, λ h, of_lintegral_ne_top h.ne⟩ - -end fin_meas_supp - -end fin_meas_supp - -/-- To prove something for an arbitrary simple function, it suffices to show -that the property holds for (multiples of) characteristic functions and is closed under -addition (of functions with disjoint support). - -It is possible to make the hypotheses in `h_add` a bit stronger, and such conditions can be added -once we need them (for example it is only necessary to consider the case where `g` is a multiple -of a characteristic function, and that this multiple doesn't appear in the image of `f`) -/ -@[elab_as_eliminator] -protected lemma induction {α γ} [measurable_space α] [add_monoid γ] {P : simple_func α γ → Prop} - (h_ind : ∀ c {s} (hs : measurable_set s), - P (simple_func.piecewise s hs (simple_func.const _ c) (simple_func.const _ 0))) - (h_add : ∀ ⦃f g : simple_func α γ⦄, disjoint (support f) (support g) → P f → P g → P (f + g)) - (f : simple_func α γ) : P f := -begin - generalize' h : f.range \ {0} = s, - rw [← finset.coe_inj, finset.coe_sdiff, finset.coe_singleton, simple_func.coe_range] at h, - revert s f h, refine finset.induction _ _, - { intros f hf, rw [finset.coe_empty, diff_eq_empty, range_subset_singleton] at hf, - convert h_ind 0 measurable_set.univ, ext x, simp [hf] }, - { intros x s hxs ih f hf, - have mx := f.measurable_set_preimage {x}, - let g := simple_func.piecewise (f ⁻¹' {x}) mx 0 f, - have Pg : P g, - { apply ih, simp only [g, simple_func.coe_piecewise, range_piecewise], - rw [image_compl_preimage, union_diff_distrib, diff_diff_comm, hf, finset.coe_insert, - insert_diff_self_of_not_mem, diff_eq_empty.mpr, set.empty_union], - { rw [set.image_subset_iff], convert set.subset_univ _, - exact preimage_const_of_mem (mem_singleton _) }, - { rwa [finset.mem_coe] }}, - convert h_add _ Pg (h_ind x mx), - { ext1 y, by_cases hy : y ∈ f ⁻¹' {x}; [simpa [hy], simp [hy]] }, - rw disjoint_iff_inf_le, - rintro y, by_cases hy : y ∈ f ⁻¹' {x}; simp [hy] } -end - -end simple_func +variables {α β γ δ : Type*} section lintegral + open simple_func + variables {m : measurable_space α} {μ ν : measure α} /-- The **lower Lebesgue integral** of a function `f` with respect to a measure `μ`. -/ @@ -2643,35 +1682,9 @@ end end lintegral -end measure_theory - -open measure_theory measure_theory.simple_func -/-- To prove something for an arbitrary measurable function into `ℝ≥0∞`, it suffices to show -that the property holds for (multiples of) characteristic functions and is closed under addition -and supremum of increasing sequences of functions. - -It is possible to make the hypotheses in the induction steps a bit stronger, and such conditions -can be added once we need them (for example in `h_add` it is only necessary to consider the sum of -a simple function with a multiple of a characteristic function and that the intersection -of their images is a subset of `{0}`. -/ -@[elab_as_eliminator] -theorem measurable.ennreal_induction {α} [measurable_space α] {P : (α → ℝ≥0∞) → Prop} - (h_ind : ∀ (c : ℝ≥0∞) ⦃s⦄, measurable_set s → P (indicator s (λ _, c))) - (h_add : ∀ ⦃f g : α → ℝ≥0∞⦄, disjoint (support f) (support g) → measurable f → measurable g → - P f → P g → P (f + g)) - (h_supr : ∀ ⦃f : ℕ → α → ℝ≥0∞⦄ (hf : ∀n, measurable (f n)) (h_mono : monotone f) - (hP : ∀ n, P (f n)), P (λ x, ⨆ n, f n x)) - ⦃f : α → ℝ≥0∞⦄ (hf : measurable f) : P f := -begin - convert h_supr (λ n, (eapprox f n).measurable) (monotone_eapprox f) _, - { ext1 x, rw [supr_eapprox_apply f hf] }, - { exact λ n, simple_func.induction (λ c s hs, h_ind c hs) - (λ f g hfg hf hg, h_add hfg f.measurable g.measurable hf hg) (eapprox f n) } -end - -namespace measure_theory +open measure_theory.simple_func -variables {α : Type*} {m m0 : measurable_space α} +variables {m m0 : measurable_space α} include m From 8130e5155d637db35907c272de9aec9dc851c03a Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 1 May 2023 05:15:31 +0000 Subject: [PATCH 0918/1291] fix(algebra/module/submodule/basic): remove `submodule_class` (#18902) This is redundant in the face of `smul_mem_class`. This also adds a missing instance. --- src/algebra/algebra/subalgebra/basic.lean | 3 +++ src/algebra/module/submodule/basic.lean | 19 ++++++------------- 2 files changed, 9 insertions(+), 13 deletions(-) diff --git a/src/algebra/algebra/subalgebra/basic.lean b/src/algebra/algebra/subalgebra/basic.lean index b0ca846d8f4d1..9cc6730b0237d 100644 --- a/src/algebra/algebra/subalgebra/basic.lean +++ b/src/algebra/algebra/subalgebra/basic.lean @@ -98,6 +98,9 @@ S.range_subset theorem smul_mem {x : A} (hx : x ∈ S) (r : R) : r • x ∈ S := (algebra.smul_def r x).symm ▸ mul_mem (S.algebra_map_mem r) hx +instance : smul_mem_class (subalgebra R A) R A := +{ smul_mem := λ S r x hx, smul_mem S hx r } + protected theorem one_mem : (1 : A) ∈ S := one_mem S protected theorem mul_mem {x y : A} (hx : x ∈ S) (hy : y ∈ S) : x * y ∈ S := mul_mem hx hy protected theorem pow_mem {x : A} (hx : x ∈ S) (n : ℕ) : x ^ n ∈ S := pow_mem hx n diff --git a/src/algebra/module/submodule/basic.lean b/src/algebra/module/submodule/basic.lean index 39aa53575d85d..6b5d2175764bb 100644 --- a/src/algebra/module/submodule/basic.lean +++ b/src/algebra/module/submodule/basic.lean @@ -35,13 +35,6 @@ variables {G : Type u''} {S : Type u'} {R : Type u} {M : Type v} {ι : Type w} set_option old_structure_cmd true -/-- `submodule_class S R M` says `S` is a type of submodules `s ≤ M`. - -Note that only `R` is marked as `out_param` since `M` is already supplied by the `set_like` class. --/ -class submodule_class (S : Type*) (R : out_param $ Type*) (M : Type*) [add_zero_class M] - [has_smul R M] [set_like S M] [add_submonoid_class S M] extends smul_mem_class S R M - /-- A submodule of a module is one which is closed under vector operations. This is a sufficient condition for the subset of vectors in the submodule to themselves form a module. -/ @@ -66,7 +59,7 @@ instance : add_submonoid_class (submodule R M) M := { zero_mem := zero_mem', add_mem := add_mem' } -instance : submodule_class (submodule R M) R M := +instance : smul_mem_class (submodule R M) R M := { smul_mem := smul_mem' } @[simp] theorem mem_to_add_submonoid (p : submodule R M) (x : M) : x ∈ p.to_add_submonoid ↔ x ∈ p := @@ -139,23 +132,23 @@ to_sub_mul_action_strict_mono.monotone end submodule -namespace submodule_class +namespace smul_mem_class variables [semiring R] [add_comm_monoid M] [module R M] {A : Type*} [set_like A M] - [add_submonoid_class A M] [hA : submodule_class A R M] (S' : A) + [add_submonoid_class A M] [hA : smul_mem_class A R M] (S' : A) include hA /-- A submodule of a `module` is a `module`. -/ -@[priority 75] -- Prefer subclasses of `module` over `submodule_class`. +@[priority 75] -- Prefer subclasses of `module` over `smul_mem_class`. instance to_module : module R S' := subtype.coe_injective.module R (add_submonoid_class.subtype S') (set_like.coe_smul S') /-- The natural `R`-linear map from a submodule of an `R`-module `M` to `M`. -/ protected def subtype : S' →ₗ[R] M := ⟨coe, λ _ _, rfl, λ _ _, rfl⟩ -@[simp] protected theorem coe_subtype : (submodule_class.subtype S' : S' → M) = coe := rfl +@[simp] protected theorem coe_subtype : (smul_mem_class.subtype S' : S' → M) = coe := rfl -end submodule_class +end smul_mem_class namespace submodule From 39478763114722f0ec7613cb2f3f7701f9b86c8d Mon Sep 17 00:00:00 2001 From: Kevin Buzzard Date: Mon, 1 May 2023 20:00:55 +0000 Subject: [PATCH 0919/1291] chore(category_theory/limits/preserves): fix typo (#18906) --- src/category_theory/functor/flat.lean | 4 ++-- src/category_theory/limits/preserves/functor_category.lean | 2 +- 2 files changed, 3 insertions(+), 3 deletions(-) diff --git a/src/category_theory/functor/flat.lean b/src/category_theory/functor/flat.lean index b8adef71d1b89..6eca7ce37093f 100644 --- a/src/category_theory/functor/flat.lean +++ b/src/category_theory/functor/flat.lean @@ -355,7 +355,7 @@ begin haveI : preserves_finite_limits F := begin apply preserves_finite_limits_of_preserves_finite_limits_of_size.{u₁}, - intros, resetI, apply preserves_limit_of_Lan_presesrves_limit + intros, resetI, apply preserves_limit_of_Lan_preserves_limit end, apply flat_of_preserves_finite_limits end⟩ @@ -371,7 +371,7 @@ def preserves_finite_limits_iff_Lan_preserves_finite_limits (F : C ⥤ D) : inv_fun := λ _, begin apply preserves_finite_limits_of_preserves_finite_limits_of_size.{u₁}, - intros, resetI, apply preserves_limit_of_Lan_presesrves_limit + intros, resetI, apply preserves_limit_of_Lan_preserves_limit end, left_inv := λ x, begin diff --git a/src/category_theory/limits/preserves/functor_category.lean b/src/category_theory/limits/preserves/functor_category.lean index 0770cdaa956a1..2db3de686ed80 100644 --- a/src/category_theory/limits/preserves/functor_category.lean +++ b/src/category_theory/limits/preserves/functor_category.lean @@ -99,7 +99,7 @@ instance whiskering_right_preserves_limits {C : Type u} [category C] /-- If `Lan F.op : (Cᵒᵖ ⥤ Type*) ⥤ (Dᵒᵖ ⥤ Type*)` preserves limits of shape `J`, so will `F`. -/ noncomputable -def preserves_limit_of_Lan_presesrves_limit {C D : Type u} [small_category C] [small_category D] +def preserves_limit_of_Lan_preserves_limit {C D : Type u} [small_category C] [small_category D] (F : C ⥤ D) (J : Type u) [small_category J] [preserves_limits_of_shape J (Lan F.op : _ ⥤ (Dᵒᵖ ⥤ Type u))] : preserves_limits_of_shape J F := From b2d7b50dc4b9fcb559e3459777e6c0c2400fe81c Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 1 May 2023 23:12:52 +0000 Subject: [PATCH 0920/1291] chore(algebra/order/to_interval_mod): Reorder arguments (#18908) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The current argument order to `to_Ixx_mod` is a bit poor: In `to_Ico a hb x`, `a` and `x` play a similar role while `b` is completely separate. This PR changes this to `to_Ico hp a b`, where the translation is as follows: * `mem_Ioo_mod a p b` → `mem_Ioo_mod p a b` (this is more consistent with `int.modeq p a b`) * `to_Ico_div a hp b` → `to_Ico_div hp a b` * `to_Ioc_div a hp b` → `to_Ioc_div hp a b` * `to_Ico_mod a hp b` → `to_Ico_mod hp a b` * `to_Ioc_mod a hp b` → `to_Ioc_mod hp a b` While it might be tempting to order`to_Ico_div hp a b` as `to_Ico_div b a hp` to match the order of the morally equivalent `(b - a) / p + a`, we opt for the simpler convention of being consistent with `mem_Ioo_mod`, as in downstream files we also find it convenient to have `b` as the last argument. Generally speaking, the variables have been renamed to * `a` → `a` (the left end) * `b` → `p` (the period) * `x` → `b` (the interval length) * `y` → `c` (additional argument similar to the left end and interval length) * `z` → `n` (the euclidean dividend) Proofs are golfed slightly. The only lemma statements that changed are `to_Ico_div_eq_of_sub_zsmul_mem_Ico`/`to_Ioc_div_eq_of_sub_zsmul_mem_Ioc` that now state `to_Ico_div hp a b = n` rather than `n = to_Ico_div hp a b` (they were always used in this new direction). --- src/algebra/order/to_interval_mod.lean | 751 +++++++----------- .../special_functions/complex/arg.lean | 8 +- .../trigonometric/angle.lean | 19 +- src/topology/instances/add_circle.lean | 42 +- 4 files changed, 334 insertions(+), 486 deletions(-) diff --git a/src/algebra/order/to_interval_mod.lean b/src/algebra/order/to_interval_mod.lean index 34a4176ebf5b6..201deaa40340c 100644 --- a/src/algebra/order/to_interval_mod.lean +++ b/src/algebra/order/to_interval_mod.lean @@ -21,12 +21,12 @@ interval. ## Main definitions -* `to_Ico_div a hb x` (where `hb : 0 < b`): The unique integer such that this multiple of `b`, - subtracted from `x`, is in `Ico a (a + b)`. -* `to_Ico_mod a hb x` (where `hb : 0 < b`): Reduce `x` to the interval `Ico a (a + b)`. -* `to_Ioc_div a hb x` (where `hb : 0 < b`): The unique integer such that this multiple of `b`, - subtracted from `x`, is in `Ioc a (a + b)`. -* `to_Ioc_mod a hb x` (where `hb : 0 < b`): Reduce `x` to the interval `Ioc a (a + b)`. +* `to_Ico_div hp a b` (where `hp : 0 < p`): The unique integer such that this multiple of `p`, + subtracted from `b`, is in `Ico a (a + p)`. +* `to_Ico_mod hp a b` (where `hp : 0 < p`): Reduce `b` to the interval `Ico a (a + p)`. +* `to_Ioc_div hp a b` (where `hp : 0 < p`): The unique integer such that this multiple of `p`, + subtracted from `b`, is in `Ioc a (a + p)`. +* `to_Ioc_mod hp a b` (where `hp : 0 < p`): Reduce `b` to the interval `Ioc a (a + p)`. -/ @@ -34,410 +34,288 @@ noncomputable theory section linear_ordered_add_comm_group -variables {α : Type*} [linear_ordered_add_comm_group α] [hα : archimedean α] +variables {α : Type*} [linear_ordered_add_comm_group α] [hα : archimedean α] {p : α} (hp : 0 < p) + {a b c : α} {n : ℤ} include hα /-- -The unique integer such that this multiple of `b`, subtracted from `x`, is in `Ico a (a + b)`. -/ -def to_Ico_div (a : α) {b : α} (hb : 0 < b) (x : α) : ℤ := -(exists_unique_sub_zsmul_mem_Ico hb x a).some +The unique integer such that this multiple of `p`, subtracted from `b`, is in `Ico a (a + p)`. -/ +def to_Ico_div (a b : α) : ℤ := (exists_unique_sub_zsmul_mem_Ico hp b a).some -lemma sub_to_Ico_div_zsmul_mem_Ico (a : α) {b : α} (hb : 0 < b) (x : α) : - x - to_Ico_div a hb x • b ∈ set.Ico a (a + b) := -(exists_unique_sub_zsmul_mem_Ico hb x a).some_spec.1 +lemma sub_to_Ico_div_zsmul_mem_Ico (a b : α) : b - to_Ico_div hp a b • p ∈ set.Ico a (a + p) := +(exists_unique_sub_zsmul_mem_Ico hp b a).some_spec.1 -lemma eq_to_Ico_div_of_sub_zsmul_mem_Ico {a b x : α} (hb : 0 < b) {y : ℤ} - (hy : x - y • b ∈ set.Ico a (a + b)) : y = to_Ico_div a hb x := -(exists_unique_sub_zsmul_mem_Ico hb x a).some_spec.2 y hy +lemma to_Ico_div_eq_of_sub_zsmul_mem_Ico (h : b - n • p ∈ set.Ico a (a + p)) : + to_Ico_div hp a b = n := +((exists_unique_sub_zsmul_mem_Ico hp b a).some_spec.2 _ h).symm /-- -The unique integer such that this multiple of `b`, subtracted from `x`, is in `Ioc a (a + b)`. -/ -def to_Ioc_div (a : α) {b : α} (hb : 0 < b) (x : α) : ℤ := -(exists_unique_sub_zsmul_mem_Ioc hb x a).some +The unique integer such that this multiple of `p`, subtracted from `b`, is in `Ioc a (a + p)`. -/ +def to_Ioc_div (a b : α) : ℤ := (exists_unique_sub_zsmul_mem_Ioc hp b a).some -lemma sub_to_Ioc_div_zsmul_mem_Ioc (a : α) {b : α} (hb : 0 < b) (x : α) : - x - to_Ioc_div a hb x • b ∈ set.Ioc a (a + b) := -(exists_unique_sub_zsmul_mem_Ioc hb x a).some_spec.1 +lemma sub_to_Ioc_div_zsmul_mem_Ioc (a b : α) : b - to_Ioc_div hp a b • p ∈ set.Ioc a (a + p) := +(exists_unique_sub_zsmul_mem_Ioc hp b a).some_spec.1 -lemma eq_to_Ioc_div_of_sub_zsmul_mem_Ioc {a b x : α} (hb : 0 < b) {y : ℤ} - (hy : x - y • b ∈ set.Ioc a (a + b)) : y = to_Ioc_div a hb x := -(exists_unique_sub_zsmul_mem_Ioc hb x a).some_spec.2 y hy +lemma to_Ioc_div_eq_of_sub_zsmul_mem_Ioc (h : b - n • p ∈ set.Ioc a (a + p)) : + to_Ioc_div hp a b = n := +((exists_unique_sub_zsmul_mem_Ioc hp b a).some_spec.2 _ h).symm -/-- Reduce `x` to the interval `Ico a (a + b)`. -/ -def to_Ico_mod (a : α) {b : α} (hb : 0 < b) (x : α) : α := x - to_Ico_div a hb x • b +/-- Reduce `b` to the interval `Ico a (a + p)`. -/ +def to_Ico_mod (a b : α) : α := b - to_Ico_div hp a b • p -/-- Reduce `x` to the interval `Ioc a (a + b)`. -/ -def to_Ioc_mod (a : α) {b : α} (hb : 0 < b) (x : α) : α := x - to_Ioc_div a hb x • b +/-- Reduce `b` to the interval `Ioc a (a + p)`. -/ +def to_Ioc_mod (a b : α) : α := b - to_Ioc_div hp a b • p -lemma to_Ico_mod_mem_Ico (a : α) {b : α} (hb : 0 < b) (x : α) : - to_Ico_mod a hb x ∈ set.Ico a (a + b) := -sub_to_Ico_div_zsmul_mem_Ico a hb x +lemma to_Ico_mod_mem_Ico (a b : α) : to_Ico_mod hp a b ∈ set.Ico a (a + p) := +sub_to_Ico_div_zsmul_mem_Ico hp a b -lemma to_Ico_mod_mem_Ico' {b : α} (hb : 0 < b) (x : α) : - to_Ico_mod 0 hb x ∈ set.Ico 0 b := -by { convert to_Ico_mod_mem_Ico 0 hb x, exact (zero_add b).symm, } +lemma to_Ico_mod_mem_Ico' (b : α) : to_Ico_mod hp 0 b ∈ set.Ico 0 p := +by { convert to_Ico_mod_mem_Ico hp 0 b, exact (zero_add p).symm, } -lemma to_Ioc_mod_mem_Ioc (a : α) {b : α} (hb : 0 < b) (x : α) : - to_Ioc_mod a hb x ∈ set.Ioc a (a + b) := -sub_to_Ioc_div_zsmul_mem_Ioc a hb x +lemma to_Ioc_mod_mem_Ioc (a b : α) : to_Ioc_mod hp a b ∈ set.Ioc a (a + p) := +sub_to_Ioc_div_zsmul_mem_Ioc hp a b -lemma left_le_to_Ico_mod (a : α) {b : α} (hb : 0 < b) (x : α) : a ≤ to_Ico_mod a hb x := -(set.mem_Ico.1 (to_Ico_mod_mem_Ico a hb x)).1 +lemma left_le_to_Ico_mod (a b : α) : a ≤ to_Ico_mod hp a b := +(set.mem_Ico.1 (to_Ico_mod_mem_Ico hp a b)).1 -lemma left_lt_to_Ioc_mod (a : α) {b : α} (hb : 0 < b) (x : α) : a < to_Ioc_mod a hb x := -(set.mem_Ioc.1 (to_Ioc_mod_mem_Ioc a hb x)).1 +lemma left_lt_to_Ioc_mod (a b : α) : a < to_Ioc_mod hp a b := +(set.mem_Ioc.1 (to_Ioc_mod_mem_Ioc hp a b)).1 -lemma to_Ico_mod_lt_right (a : α) {b : α} (hb : 0 < b) (x : α) : to_Ico_mod a hb x < a + b := -(set.mem_Ico.1 (to_Ico_mod_mem_Ico a hb x)).2 +lemma to_Ico_mod_lt_right (a b : α) : to_Ico_mod hp a b < a + p := +(set.mem_Ico.1 (to_Ico_mod_mem_Ico hp a b)).2 -lemma to_Ioc_mod_le_right (a : α) {b : α} (hb : 0 < b) (x : α) : to_Ioc_mod a hb x ≤ a + b := -(set.mem_Ioc.1 (to_Ioc_mod_mem_Ioc a hb x)).2 +lemma to_Ioc_mod_le_right (a b : α) : to_Ioc_mod hp a b ≤ a + p := +(set.mem_Ioc.1 (to_Ioc_mod_mem_Ioc hp a b)).2 -@[simp] lemma self_sub_to_Ico_div_zsmul (a : α) {b : α} (hb : 0 < b) (x : α) : - x - to_Ico_div a hb x • b = to_Ico_mod a hb x := +@[simp] lemma self_sub_to_Ico_div_zsmul (a b : α) : b - to_Ico_div hp a b • p = to_Ico_mod hp a b := rfl -@[simp] lemma self_sub_to_Ioc_div_zsmul (a : α) {b : α} (hb : 0 < b) (x : α) : - x - to_Ioc_div a hb x • b = to_Ioc_mod a hb x := +@[simp] lemma self_sub_to_Ioc_div_zsmul (a b : α) : b - to_Ioc_div hp a b • p = to_Ioc_mod hp a b := rfl -@[simp] lemma to_Ico_div_zsmul_sub_self (a : α) {b : α} (hb : 0 < b) (x : α) : - to_Ico_div a hb x • b - x = -to_Ico_mod a hb x := +@[simp] lemma to_Ico_div_zsmul_sub_self (a b : α) : + to_Ico_div hp a b • p - b = -to_Ico_mod hp a b := by rw [to_Ico_mod, neg_sub] -@[simp] lemma to_Ioc_div_zsmul_sub_self (a : α) {b : α} (hb : 0 < b) (x : α) : - to_Ioc_div a hb x • b - x = -to_Ioc_mod a hb x := +@[simp] lemma to_Ioc_div_zsmul_sub_self (a b : α) : + to_Ioc_div hp a b • p - b = -to_Ioc_mod hp a b := by rw [to_Ioc_mod, neg_sub] -@[simp] lemma to_Ico_mod_sub_self (a : α) {b : α} (hb : 0 < b) (x : α) : - to_Ico_mod a hb x - x = -to_Ico_div a hb x • b := +@[simp] lemma to_Ico_mod_sub_self (a b : α) : to_Ico_mod hp a b - b = -to_Ico_div hp a b • p := by rw [to_Ico_mod, sub_sub_cancel_left, neg_smul] -@[simp] lemma to_Ioc_mod_sub_self (a : α) {b : α} (hb : 0 < b) (x : α) : - to_Ioc_mod a hb x - x = -to_Ioc_div a hb x • b := +@[simp] lemma to_Ioc_mod_sub_self (a b : α) : to_Ioc_mod hp a b - b = -to_Ioc_div hp a b • p := by rw [to_Ioc_mod, sub_sub_cancel_left, neg_smul] -@[simp] lemma self_sub_to_Ico_mod (a : α) {b : α} (hb : 0 < b) (x : α) : - x - to_Ico_mod a hb x = to_Ico_div a hb x • b := +@[simp] lemma self_sub_to_Ico_mod (a b : α) : b - to_Ico_mod hp a b = to_Ico_div hp a b • p := by rw [to_Ico_mod, sub_sub_cancel] -@[simp] lemma self_sub_to_Ioc_mod (a : α) {b : α} (hb : 0 < b) (x : α) : - x - to_Ioc_mod a hb x = to_Ioc_div a hb x • b := +@[simp] lemma self_sub_to_Ioc_mod (a b : α) : b - to_Ioc_mod hp a b = to_Ioc_div hp a b • p := by rw [to_Ioc_mod, sub_sub_cancel] -@[simp] lemma to_Ico_mod_add_to_Ico_div_zsmul (a : α) {b : α} (hb : 0 < b) (x : α) : - to_Ico_mod a hb x + to_Ico_div a hb x • b = x := +@[simp] lemma to_Ico_mod_add_to_Ico_div_zsmul (a b : α) : + to_Ico_mod hp a b + to_Ico_div hp a b • p = b := by rw [to_Ico_mod, sub_add_cancel] -@[simp] lemma to_Ioc_mod_add_to_Ioc_div_zsmul (a : α) {b : α} (hb : 0 < b) (x : α) : - to_Ioc_mod a hb x + to_Ioc_div a hb x • b = x := +@[simp] lemma to_Ioc_mod_add_to_Ioc_div_zsmul (a b : α) : + to_Ioc_mod hp a b + to_Ioc_div hp a b • p = b := by rw [to_Ioc_mod, sub_add_cancel] -@[simp] lemma to_Ico_div_zsmul_sub_to_Ico_mod (a : α) {b : α} (hb : 0 < b) (x : α) : - to_Ico_div a hb x • b + to_Ico_mod a hb x = x := +@[simp] lemma to_Ico_div_zsmul_sub_to_Ico_mod (a b : α) : + to_Ico_div hp a b • p + to_Ico_mod hp a b = b := by rw [add_comm, to_Ico_mod_add_to_Ico_div_zsmul] -@[simp] lemma to_Ioc_div_zsmul_sub_to_Ioc_mod (a : α) {b : α} (hb : 0 < b) (x : α) : - to_Ioc_div a hb x • b + to_Ioc_mod a hb x = x := +@[simp] lemma to_Ioc_div_zsmul_sub_to_Ioc_mod (a b : α) : + to_Ioc_div hp a b • p + to_Ioc_mod hp a b = b := by rw [add_comm, to_Ioc_mod_add_to_Ioc_div_zsmul] -lemma to_Ico_mod_eq_iff {a b x y : α} (hb : 0 < b) : - to_Ico_mod a hb x = y ↔ y ∈ set.Ico a (a + b) ∧ ∃ z : ℤ, x = y + z • b := +lemma to_Ico_mod_eq_iff : to_Ico_mod hp a b = c ↔ c ∈ set.Ico a (a + p) ∧ ∃ z : ℤ, b = c + z • p := begin - refine ⟨λ h, ⟨h ▸ to_Ico_mod_mem_Ico a hb x, - to_Ico_div a hb x, - h ▸ (to_Ico_mod_add_to_Ico_div_zsmul _ _ _).symm⟩, - λ h, _⟩, - rcases h with ⟨hy, z, hz⟩, - rw ←sub_eq_iff_eq_add at hz, - subst hz, - rw eq_to_Ico_div_of_sub_zsmul_mem_Ico hb hy, - refl + refine ⟨λ h, ⟨h ▸ to_Ico_mod_mem_Ico hp a b, to_Ico_div hp a b, + h ▸ (to_Ico_mod_add_to_Ico_div_zsmul _ _ _).symm⟩, _⟩, + simp_rw ←@sub_eq_iff_eq_add, + rintro ⟨hc, n, rfl⟩, + rw [←to_Ico_div_eq_of_sub_zsmul_mem_Ico hp hc, to_Ico_mod], end -lemma to_Ioc_mod_eq_iff {a b x y : α} (hb : 0 < b) : - to_Ioc_mod a hb x = y ↔ y ∈ set.Ioc a (a + b) ∧ ∃ z : ℤ, x = y + z • b := +lemma to_Ioc_mod_eq_iff : to_Ioc_mod hp a b = c ↔ c ∈ set.Ioc a (a + p) ∧ ∃ z : ℤ, b = c + z • p := begin - refine ⟨λ h, ⟨h ▸ to_Ioc_mod_mem_Ioc a hb x, - to_Ioc_div a hb x, - h ▸ (to_Ioc_mod_add_to_Ioc_div_zsmul _ hb _).symm⟩, - λ h, _⟩, - rcases h with ⟨hy, z, hz⟩, - rw ←sub_eq_iff_eq_add at hz, - subst hz, - rw eq_to_Ioc_div_of_sub_zsmul_mem_Ioc hb hy, - refl + refine ⟨λ h, ⟨h ▸ to_Ioc_mod_mem_Ioc hp a b, to_Ioc_div hp a b, + h ▸ (to_Ioc_mod_add_to_Ioc_div_zsmul hp _ _).symm⟩, _⟩, + simp_rw ←@sub_eq_iff_eq_add, + rintro ⟨hc, n, rfl⟩, + rw [←to_Ioc_div_eq_of_sub_zsmul_mem_Ioc hp hc, to_Ioc_mod], end -@[simp] lemma to_Ico_div_apply_left (a : α) {b : α} (hb : 0 < b) : to_Ico_div a hb a = 0 := -begin - refine (eq_to_Ico_div_of_sub_zsmul_mem_Ico hb _).symm, - simp [hb] -end +@[simp] lemma to_Ico_div_apply_left (a : α) : to_Ico_div hp a a = 0 := +to_Ico_div_eq_of_sub_zsmul_mem_Ico hp $ by simp [hp] -@[simp] lemma to_Ioc_div_apply_left (a : α) {b : α} (hb : 0 < b) : to_Ioc_div a hb a = -1 := -begin - refine (eq_to_Ioc_div_of_sub_zsmul_mem_Ioc hb _).symm, - simp [hb], -end +@[simp] lemma to_Ioc_div_apply_left (a : α) : to_Ioc_div hp a a = -1 := +to_Ioc_div_eq_of_sub_zsmul_mem_Ioc hp $ by simp [hp] -@[simp] lemma to_Ico_mod_apply_left (a : α) {b : α} (hb : 0 < b) : to_Ico_mod a hb a = a := -begin - rw [to_Ico_mod_eq_iff hb, set.left_mem_Ico], - refine ⟨lt_add_of_pos_right _ hb, 0, _⟩, - simp -end +@[simp] lemma to_Ico_mod_apply_left (a : α) : to_Ico_mod hp a a = a := +by { rw [to_Ico_mod_eq_iff hp, set.left_mem_Ico], exact ⟨lt_add_of_pos_right _ hp, 0, by simp⟩ } -@[simp] lemma to_Ioc_mod_apply_left (a : α) {b : α} (hb : 0 < b) : to_Ioc_mod a hb a = a + b := -begin - rw [to_Ioc_mod_eq_iff hb, set.right_mem_Ioc], - refine ⟨lt_add_of_pos_right _ hb, -1, _⟩, - simp -end +@[simp] lemma to_Ioc_mod_apply_left (a : α) : to_Ioc_mod hp a a = a + p := +by { rw [to_Ioc_mod_eq_iff hp, set.right_mem_Ioc], exact ⟨lt_add_of_pos_right _ hp, -1, by simp⟩ } -lemma to_Ico_div_apply_right (a : α) {b : α} (hb : 0 < b) : - to_Ico_div a hb (a + b) = 1 := -begin - refine (eq_to_Ico_div_of_sub_zsmul_mem_Ico hb _).symm, - simp [hb] -end +lemma to_Ico_div_apply_right (a : α) : to_Ico_div hp a (a + p) = 1 := +to_Ico_div_eq_of_sub_zsmul_mem_Ico hp $ by simp [hp] -lemma to_Ioc_div_apply_right (a : α) {b : α} (hb : 0 < b) : - to_Ioc_div a hb (a + b) = 0 := -begin - refine (eq_to_Ioc_div_of_sub_zsmul_mem_Ioc hb _).symm, - simp [hb] -end +lemma to_Ioc_div_apply_right (a : α) : to_Ioc_div hp a (a + p) = 0 := +to_Ioc_div_eq_of_sub_zsmul_mem_Ioc hp $ by simp [hp] -lemma to_Ico_mod_apply_right (a : α) {b : α} (hb : 0 < b) : to_Ico_mod a hb (a + b) = a := -begin - rw [to_Ico_mod_eq_iff hb, set.left_mem_Ico], - refine ⟨lt_add_of_pos_right _ hb, 1, _⟩, - simp -end +lemma to_Ico_mod_apply_right (a : α) : to_Ico_mod hp a (a + p) = a := +by { rw [to_Ico_mod_eq_iff hp, set.left_mem_Ico], exact ⟨lt_add_of_pos_right _ hp, 1, by simp⟩ } -lemma to_Ioc_mod_apply_right (a : α) {b : α} (hb : 0 < b) : - to_Ioc_mod a hb (a + b) = a + b := -begin - rw [to_Ioc_mod_eq_iff hb, set.right_mem_Ioc], - refine ⟨lt_add_of_pos_right _ hb, 0, _⟩, - simp -end +lemma to_Ioc_mod_apply_right (a : α) : to_Ioc_mod hp a (a + p) = a + p := +by { rw [to_Ioc_mod_eq_iff hp, set.right_mem_Ioc], exact ⟨lt_add_of_pos_right _ hp, 0, by simp⟩ } -@[simp] lemma to_Ico_div_add_zsmul (a : α) {b : α} (hb : 0 < b) (x : α) (m : ℤ) : - to_Ico_div a hb (x + m • b) = to_Ico_div a hb x + m := -begin - refine (eq_to_Ico_div_of_sub_zsmul_mem_Ico hb _).symm, - convert sub_to_Ico_div_zsmul_mem_Ico a hb x using 1, - simp [add_smul], -end +@[simp] lemma to_Ico_div_add_zsmul (a b : α) (m : ℤ) : + to_Ico_div hp a (b + m • p) = to_Ico_div hp a b + m := +to_Ico_div_eq_of_sub_zsmul_mem_Ico hp $ + by simpa only [add_smul, add_sub_add_right_eq_sub] using sub_to_Ico_div_zsmul_mem_Ico hp a b -@[simp] lemma to_Ioc_div_add_zsmul (a : α) {b : α} (hb : 0 < b) (x : α) (m : ℤ) : - to_Ioc_div a hb (x + m • b) = to_Ioc_div a hb x + m := -begin - refine (eq_to_Ioc_div_of_sub_zsmul_mem_Ioc hb _).symm, - convert sub_to_Ioc_div_zsmul_mem_Ioc a hb x using 1, - simp [add_smul] -end +@[simp] lemma to_Ioc_div_add_zsmul (a b : α) (m : ℤ) : + to_Ioc_div hp a (b + m • p) = to_Ioc_div hp a b + m := +to_Ioc_div_eq_of_sub_zsmul_mem_Ioc hp $ + by simpa only [add_smul, add_sub_add_right_eq_sub] using sub_to_Ioc_div_zsmul_mem_Ioc hp a b -@[simp] lemma to_Ico_div_zsmul_add (a : α) {b : α} (hb : 0 < b) (x : α) (m : ℤ) : - to_Ico_div a hb (m • b + x) = m + to_Ico_div a hb x := +@[simp] lemma to_Ico_div_zsmul_add (a b : α) (m : ℤ) : + to_Ico_div hp a (m • p + b) = m + to_Ico_div hp a b := by rw [add_comm, to_Ico_div_add_zsmul, add_comm] -@[simp] lemma to_Ioc_div_zsmul_add (a : α) {b : α} (hb : 0 < b) (x : α) (m : ℤ) : - to_Ioc_div a hb (m • b + x) = to_Ioc_div a hb x + m := +@[simp] lemma to_Ioc_div_zsmul_add (a b : α) (m : ℤ) : + to_Ioc_div hp a (m • p + b) = to_Ioc_div hp a b + m := by rw [add_comm, to_Ioc_div_add_zsmul, add_comm] -@[simp] lemma to_Ico_div_sub_zsmul (a : α) {b : α} (hb : 0 < b) (x : α) (m : ℤ) : - to_Ico_div a hb (x - m • b) = to_Ico_div a hb x - m := +@[simp] lemma to_Ico_div_sub_zsmul (a b : α) (m : ℤ) : + to_Ico_div hp a (b - m • p) = to_Ico_div hp a b - m := by rw [sub_eq_add_neg, ←neg_smul, to_Ico_div_add_zsmul, sub_eq_add_neg] -@[simp] lemma to_Ioc_div_sub_zsmul (a : α) {b : α} (hb : 0 < b) (x : α) (m : ℤ) : - to_Ioc_div a hb (x - m • b) = to_Ioc_div a hb x - m := +@[simp] lemma to_Ioc_div_sub_zsmul (a b : α) (m : ℤ) : + to_Ioc_div hp a (b - m • p) = to_Ioc_div hp a b - m := by rw [sub_eq_add_neg, ←neg_smul, to_Ioc_div_add_zsmul, sub_eq_add_neg] -@[simp] lemma to_Ico_div_add_right (a : α) {b : α} (hb : 0 < b) (x : α) : - to_Ico_div a hb (x + b) = to_Ico_div a hb x + 1 := -begin - convert to_Ico_div_add_zsmul a hb x 1, - exact (one_zsmul _).symm -end +@[simp] lemma to_Ico_div_add_right (a b : α) : to_Ico_div hp a (b + p) = to_Ico_div hp a b + 1 := +by simpa only [one_zsmul] using to_Ico_div_add_zsmul hp a b 1 -@[simp] lemma to_Ioc_div_add_right (a : α) {b : α} (hb : 0 < b) (x : α) : - to_Ioc_div a hb (x + b) = to_Ioc_div a hb x + 1 := -begin - convert to_Ioc_div_add_zsmul a hb x 1, - exact (one_zsmul _).symm -end +@[simp] lemma to_Ioc_div_add_right (a b : α) : to_Ioc_div hp a (b + p) = to_Ioc_div hp a b + 1 := +by simpa only [one_zsmul] using to_Ioc_div_add_zsmul hp a b 1 -@[simp] lemma to_Ico_div_add_left (a : α) {b : α} (hb : 0 < b) (x : α) : - to_Ico_div a hb (b + x) = to_Ico_div a hb x + 1 := +@[simp] lemma to_Ico_div_add_left (a b : α) : to_Ico_div hp a (p + b) = to_Ico_div hp a b + 1 := by rw [add_comm, to_Ico_div_add_right] -@[simp] lemma to_Ioc_div_add_left (a : α) {b : α} (hb : 0 < b) (x : α) : - to_Ioc_div a hb (b + x) = to_Ioc_div a hb x + 1 := +@[simp] lemma to_Ioc_div_add_left (a b : α) : to_Ioc_div hp a (p + b) = to_Ioc_div hp a b + 1 := by rw [add_comm, to_Ioc_div_add_right] -@[simp] lemma to_Ico_div_sub (a : α) {b : α} (hb : 0 < b) (x : α) : - to_Ico_div a hb (x - b) = to_Ico_div a hb x - 1 := -begin - convert to_Ico_div_sub_zsmul a hb x 1, - exact (one_zsmul _).symm -end +@[simp] lemma to_Ico_div_sub (a b : α) : to_Ico_div hp a (b - p) = to_Ico_div hp a b - 1 := +by simpa only [one_zsmul] using to_Ico_div_sub_zsmul hp a b 1 -@[simp] lemma to_Ioc_div_sub (a : α) {b : α} (hb : 0 < b) (x : α) : - to_Ioc_div a hb (x - b) = to_Ioc_div a hb x - 1 := -begin - convert to_Ioc_div_sub_zsmul a hb x 1, - exact (one_zsmul _).symm -end +@[simp] lemma to_Ioc_div_sub (a b : α) : to_Ioc_div hp a (b - p) = to_Ioc_div hp a b - 1 := +by simpa only [one_zsmul] using to_Ioc_div_sub_zsmul hp a b 1 -lemma to_Ico_div_sub' (a : α) {b : α} (hb : 0 < b) (x y : α) : - to_Ico_div a hb (x - y) = to_Ico_div (a + y) hb x := +lemma to_Ico_div_sub' (a b c : α) : to_Ico_div hp a (b - c) = to_Ico_div hp (a + c) b := begin - rw eq_comm, - apply eq_to_Ico_div_of_sub_zsmul_mem_Ico, + apply to_Ico_div_eq_of_sub_zsmul_mem_Ico, rw [←sub_right_comm, set.sub_mem_Ico_iff_left, add_right_comm], - exact sub_to_Ico_div_zsmul_mem_Ico (a + y) hb x, + exact sub_to_Ico_div_zsmul_mem_Ico hp (a + c) b, end -lemma to_Ioc_div_sub' (a : α) {b : α} (hb : 0 < b) (x y : α) : - to_Ioc_div a hb (x - y) = to_Ioc_div (a + y) hb x := +lemma to_Ioc_div_sub' (a b c : α) : to_Ioc_div hp a (b - c) = to_Ioc_div hp (a + c) b := begin - rw eq_comm, - apply eq_to_Ioc_div_of_sub_zsmul_mem_Ioc, + apply to_Ioc_div_eq_of_sub_zsmul_mem_Ioc, rw [←sub_right_comm, set.sub_mem_Ioc_iff_left, add_right_comm], - exact sub_to_Ioc_div_zsmul_mem_Ioc (a + y) hb x, + exact sub_to_Ioc_div_zsmul_mem_Ioc hp (a + c) b, end -lemma to_Ico_div_add_right' (a : α) {b : α} (hb : 0 < b) (x y : α) : - to_Ico_div a hb (x + y) = to_Ico_div (a - y) hb x := +lemma to_Ico_div_add_right' (a b c : α) : to_Ico_div hp a (b + c) = to_Ico_div hp (a - c) b := by rw [←sub_neg_eq_add, to_Ico_div_sub', sub_eq_add_neg] -lemma to_Ioc_div_add_right' (a : α) {b : α} (hb : 0 < b) (x y : α) : - to_Ioc_div a hb (x + y) = to_Ioc_div (a - y) hb x := +lemma to_Ioc_div_add_right' (a b c : α) : to_Ioc_div hp a (b + c) = to_Ioc_div hp (a - c) b := by rw [←sub_neg_eq_add, to_Ioc_div_sub', sub_eq_add_neg] -lemma to_Ico_div_neg (a : α) {b : α} (hb : 0 < b) (x : α) : - to_Ico_div a hb (-x) = -(to_Ioc_div (-a) hb x + 1) := +lemma to_Ico_div_neg (a b : α) : to_Ico_div hp a (-b) = -(to_Ioc_div hp (-a) b + 1) := begin - suffices : to_Ico_div a hb (-x) = -(to_Ioc_div (-(a + b)) hb x), + suffices : to_Ico_div hp a (-b) = -(to_Ioc_div hp (-(a + p)) b), { rwa [neg_add, ←sub_eq_add_neg, ←to_Ioc_div_add_right', to_Ioc_div_add_right] at this }, - rw [← neg_eq_iff_eq_neg], - apply eq_to_Ioc_div_of_sub_zsmul_mem_Ioc, - obtain ⟨hc, ho⟩ := sub_to_Ico_div_zsmul_mem_Ico a hb (-x), - rw [←neg_lt_neg_iff, neg_sub' (-x), neg_neg, ←neg_smul] at ho, - rw [←neg_le_neg_iff, neg_sub' (-x), neg_neg, ←neg_smul] at hc, + rw [← neg_eq_iff_eq_neg, eq_comm], + apply to_Ioc_div_eq_of_sub_zsmul_mem_Ioc, + obtain ⟨hc, ho⟩ := sub_to_Ico_div_zsmul_mem_Ico hp a (-b), + rw [←neg_lt_neg_iff, neg_sub' (-b), neg_neg, ←neg_smul] at ho, + rw [←neg_le_neg_iff, neg_sub' (-b), neg_neg, ←neg_smul] at hc, refine ⟨ho, hc.trans_eq _⟩, rw [neg_add, neg_add_cancel_right], end -lemma to_Ioc_div_neg (a : α) {b : α} (hb : 0 < b) (x : α) : - to_Ioc_div a hb (-x) = -(to_Ico_div (-a) hb x + 1) := -by rw [←neg_neg x, to_Ico_div_neg, neg_neg, neg_neg, neg_add', neg_neg, add_sub_cancel] +lemma to_Ioc_div_neg (a b : α) : to_Ioc_div hp a (-b) = -(to_Ico_div hp (-a) b + 1) := +by rw [←neg_neg b, to_Ico_div_neg, neg_neg, neg_neg, neg_add', neg_neg, add_sub_cancel] -@[simp] lemma to_Ico_mod_add_zsmul (a : α) {b : α} (hb : 0 < b) (x : α) (m : ℤ) : - to_Ico_mod a hb (x + m • b) = to_Ico_mod a hb x := -begin - rw [to_Ico_mod, to_Ico_div_add_zsmul, to_Ico_mod, add_smul], - abel -end +@[simp] lemma to_Ico_mod_add_zsmul (a b : α) (m : ℤ) : + to_Ico_mod hp a (b + m • p) = to_Ico_mod hp a b := +by { rw [to_Ico_mod, to_Ico_div_add_zsmul, to_Ico_mod, add_smul], abel } -@[simp] lemma to_Ioc_mod_add_zsmul (a : α) {b : α} (hb : 0 < b) (x : α) (m : ℤ) : - to_Ioc_mod a hb (x + m • b) = to_Ioc_mod a hb x := -begin - rw [to_Ioc_mod, to_Ioc_div_add_zsmul, to_Ioc_mod, add_smul], - abel -end +@[simp] lemma to_Ioc_mod_add_zsmul (a b : α) (m : ℤ) : + to_Ioc_mod hp a (b + m • p) = to_Ioc_mod hp a b := +by { rw [to_Ioc_mod, to_Ioc_div_add_zsmul, to_Ioc_mod, add_smul], abel } -@[simp] lemma to_Ico_mod_zsmul_add (a : α) {b : α} (hb : 0 < b) (x : α) (m : ℤ) : - to_Ico_mod a hb (m • b + x) = to_Ico_mod a hb x := +@[simp] lemma to_Ico_mod_zsmul_add (a b : α) (m : ℤ) : + to_Ico_mod hp a (m • p + b) = to_Ico_mod hp a b := by rw [add_comm, to_Ico_mod_add_zsmul] -@[simp] lemma to_Ioc_mod_zsmul_add (a : α) {b : α} (hb : 0 < b) (x : α) (m : ℤ) : - to_Ioc_mod a hb (m • b + x) = to_Ioc_mod a hb x := +@[simp] lemma to_Ioc_mod_zsmul_add (a b : α) (m : ℤ) : + to_Ioc_mod hp a (m • p + b) = to_Ioc_mod hp a b := by rw [add_comm, to_Ioc_mod_add_zsmul] -@[simp] lemma to_Ico_mod_sub_zsmul (a : α) {b : α} (hb : 0 < b) (x : α) (m : ℤ) : - to_Ico_mod a hb (x - m • b) = to_Ico_mod a hb x := +@[simp] lemma to_Ico_mod_sub_zsmul (a b : α) (m : ℤ) : + to_Ico_mod hp a (b - m • p) = to_Ico_mod hp a b := by rw [sub_eq_add_neg, ←neg_smul, to_Ico_mod_add_zsmul] -@[simp] lemma to_Ioc_mod_sub_zsmul (a : α) {b : α} (hb : 0 < b) (x : α) (m : ℤ) : - to_Ioc_mod a hb (x - m • b) = to_Ioc_mod a hb x := +@[simp] lemma to_Ioc_mod_sub_zsmul (a b : α) (m : ℤ) : + to_Ioc_mod hp a (b - m • p) = to_Ioc_mod hp a b := by rw [sub_eq_add_neg, ←neg_smul, to_Ioc_mod_add_zsmul] -@[simp] lemma to_Ico_mod_add_right (a : α) {b : α} (hb : 0 < b) (x : α) : - to_Ico_mod a hb (x + b) = to_Ico_mod a hb x := -begin - convert to_Ico_mod_add_zsmul a hb x 1, - exact (one_zsmul _).symm -end +@[simp] lemma to_Ico_mod_add_right (a b : α) : to_Ico_mod hp a (b + p) = to_Ico_mod hp a b := +by simpa only [one_zsmul] using to_Ico_mod_add_zsmul hp a b 1 -@[simp] lemma to_Ioc_mod_add_right (a : α) {b : α} (hb : 0 < b) (x : α) : - to_Ioc_mod a hb (x + b) = to_Ioc_mod a hb x := -begin - convert to_Ioc_mod_add_zsmul a hb x 1, - exact (one_zsmul _).symm -end +@[simp] lemma to_Ioc_mod_add_right (a b : α) : to_Ioc_mod hp a (b + p) = to_Ioc_mod hp a b := +by simpa only [one_zsmul] using to_Ioc_mod_add_zsmul hp a b 1 -@[simp] lemma to_Ico_mod_add_left (a : α) {b : α} (hb : 0 < b) (x : α) : - to_Ico_mod a hb (b + x) = to_Ico_mod a hb x := +@[simp] lemma to_Ico_mod_add_left (a b : α) : to_Ico_mod hp a (p + b) = to_Ico_mod hp a b := by rw [add_comm, to_Ico_mod_add_right] -@[simp] lemma to_Ioc_mod_add_left (a : α) {b : α} (hb : 0 < b) (x : α) : - to_Ioc_mod a hb (b + x) = to_Ioc_mod a hb x := +@[simp] lemma to_Ioc_mod_add_left (a b : α) : to_Ioc_mod hp a (p + b) = to_Ioc_mod hp a b := by rw [add_comm, to_Ioc_mod_add_right] -@[simp] lemma to_Ico_mod_sub (a : α) {b : α} (hb : 0 < b) (x : α) : - to_Ico_mod a hb (x - b) = to_Ico_mod a hb x := -begin - convert to_Ico_mod_sub_zsmul a hb x 1, - exact (one_zsmul _).symm -end +@[simp] lemma to_Ico_mod_sub (a b : α) : to_Ico_mod hp a (b - p) = to_Ico_mod hp a b := +by simpa only [one_zsmul] using to_Ico_mod_sub_zsmul hp a b 1 -@[simp] lemma to_Ioc_mod_sub (a : α) {b : α} (hb : 0 < b) (x : α) : - to_Ioc_mod a hb (x - b) = to_Ioc_mod a hb x := -begin - convert to_Ioc_mod_sub_zsmul a hb x 1, - exact (one_zsmul _).symm -end +@[simp] lemma to_Ioc_mod_sub (a b : α) : to_Ioc_mod hp a (b - p) = to_Ioc_mod hp a b := +by simpa only [one_zsmul] using to_Ioc_mod_sub_zsmul hp a b 1 -lemma to_Ico_mod_sub' (a : α) {b : α} (hb : 0 < b) (x y : α) : - to_Ico_mod a hb (x - y) = to_Ico_mod (a + y) hb x - y := +lemma to_Ico_mod_sub' (a b c : α) : to_Ico_mod hp a (b - c) = to_Ico_mod hp (a + c) b - c := by simp_rw [to_Ico_mod, to_Ico_div_sub', sub_right_comm] -lemma to_Ioc_mod_sub' (a : α) {b : α} (hb : 0 < b) (x y : α) : - to_Ioc_mod a hb (x - y) = to_Ioc_mod (a + y) hb x - y := +lemma to_Ioc_mod_sub' (a b c : α) : to_Ioc_mod hp a (b - c) = to_Ioc_mod hp (a + c) b - c := by simp_rw [to_Ioc_mod, to_Ioc_div_sub', sub_right_comm] -lemma to_Ico_mod_add_right' (a : α) {b : α} (hb : 0 < b) (x y : α) : - to_Ico_mod a hb (x + y) = to_Ico_mod (a - y) hb x + y := +lemma to_Ico_mod_add_right' (a b c : α) : to_Ico_mod hp a (b + c) = to_Ico_mod hp (a - c) b + c := by simp_rw [to_Ico_mod, to_Ico_div_add_right', sub_add_eq_add_sub] -lemma to_Ioc_mod_add_right' (a : α) {b : α} (hb : 0 < b) (x y : α) : - to_Ioc_mod a hb (x + y) = to_Ioc_mod (a - y) hb x + y := +lemma to_Ioc_mod_add_right' (a b c : α) : to_Ioc_mod hp a (b + c) = to_Ioc_mod hp (a - c) b + c := by simp_rw [to_Ioc_mod, to_Ioc_div_add_right', sub_add_eq_add_sub] -lemma to_Ico_mod_neg (a : α) {b : α} (hb : 0 < b) (x : α) : - to_Ico_mod a hb (-x) = b - to_Ioc_mod (-a) hb x := -begin - simp_rw [to_Ico_mod, to_Ioc_mod, to_Ico_div_neg, neg_smul, add_smul], - abel, -end +lemma to_Ico_mod_neg (a b : α) : to_Ico_mod hp a (-b) = p - to_Ioc_mod hp (-a) b := +by { simp_rw [to_Ico_mod, to_Ioc_mod, to_Ico_div_neg, neg_smul, add_smul], abel } -lemma to_Ioc_mod_neg (a : α) {b : α} (hb : 0 < b) (x : α) : - to_Ioc_mod a hb (-x) = b - to_Ico_mod (-a) hb x := -begin - simp_rw [to_Ioc_mod, to_Ico_mod, to_Ioc_div_neg, neg_smul, add_smul], - abel, -end +lemma to_Ioc_mod_neg (a b : α) : to_Ioc_mod hp a (-b) = p - to_Ico_mod hp (-a) b := +by { simp_rw [to_Ioc_mod, to_Ico_mod, to_Ioc_div_neg, neg_smul, add_smul], abel } -lemma to_Ico_mod_eq_to_Ico_mod (a : α) {b x y : α} (hb : 0 < b) : - to_Ico_mod a hb x = to_Ico_mod a hb y ↔ ∃ z : ℤ, y - x = z • b := +lemma to_Ico_mod_eq_to_Ico_mod : to_Ico_mod hp a b = to_Ico_mod hp a c ↔ ∃ n : ℤ, c - b = n • p := begin - refine ⟨λ h, ⟨to_Ico_div a hb y - to_Ico_div a hb x, _⟩, λ h, _⟩, - { conv_lhs { rw [←to_Ico_mod_add_to_Ico_div_zsmul a hb x, - ←to_Ico_mod_add_to_Ico_div_zsmul a hb y] }, + refine ⟨λ h, ⟨to_Ico_div hp a c - to_Ico_div hp a b, _⟩, λ h, _⟩, + { conv_lhs { rw [←to_Ico_mod_add_to_Ico_div_zsmul hp a b, + ←to_Ico_mod_add_to_Ico_div_zsmul hp a c] }, rw [h, sub_smul], abel }, { rcases h with ⟨z, hz⟩, @@ -445,12 +323,11 @@ begin rw [hz, to_Ico_mod_zsmul_add] } end -lemma to_Ioc_mod_eq_to_Ioc_mod (a : α) {b x y : α} (hb : 0 < b) : - to_Ioc_mod a hb x = to_Ioc_mod a hb y ↔ ∃ z : ℤ, y - x = z • b := +lemma to_Ioc_mod_eq_to_Ioc_mod : to_Ioc_mod hp a b = to_Ioc_mod hp a c ↔ ∃ n : ℤ, c - b = n • p := begin - refine ⟨λ h, ⟨to_Ioc_div a hb y - to_Ioc_div a hb x, _⟩, λ h, _⟩, - { conv_lhs { rw [←to_Ioc_mod_add_to_Ioc_div_zsmul a hb x, - ←to_Ioc_mod_add_to_Ioc_div_zsmul a hb y] }, + refine ⟨λ h, ⟨to_Ioc_div hp a c - to_Ioc_div hp a b, _⟩, λ h, _⟩, + { conv_lhs { rw [←to_Ioc_mod_add_to_Ioc_div_zsmul hp a b, + ←to_Ioc_mod_add_to_Ioc_div_zsmul hp a c] }, rw [h, sub_smul], abel }, { rcases h with ⟨z, hz⟩, @@ -461,278 +338,250 @@ end /-! ### Links between the `Ico` and `Ioc` variants applied to the same element -/ section Ico_Ioc - -variables (a : α) {b : α} (hb : 0 < b) (x : α) +variables (a b) omit hα -/-- `mem_Ioo_mod a b x` means that `x` lies in the open interval `(a, a + b)` modulo `b`. -Equivalently (as shown below), `x` is not congruent to `a` modulo `b`, or `to_Ico_mod a hb` agrees -with `to_Ioc_mod a hb` at `x`, or `to_Ico_div a hb` agrees with `to_Ioc_div a hb` at `x`. -/ -def mem_Ioo_mod (b x : α) : Prop := ∃ z : ℤ, x - z • b ∈ set.Ioo a (a + b) +/-- `mem_Ioo_mod p a b` means that `b` lies in the open interval `(a, a + p)` modulo `p`. +Equivalently (as shown below), `b` is not congruent to `a` modulo `p`, or `to_Ico_mod hp a` agrees +with `to_Ioc_mod hp a` at `b`, or `to_Ico_div hp a` agrees with `to_Ioc_div hp a` at `b`. -/ +def mem_Ioo_mod (p a b : α) : Prop := ∃ z : ℤ, b - z • p ∈ set.Ioo a (a + p) include hα lemma tfae_mem_Ioo_mod : - tfae [mem_Ioo_mod a b x, - to_Ico_mod a hb x = to_Ioc_mod a hb x, - to_Ico_mod a hb x + b ≠ to_Ioc_mod a hb x, - to_Ico_mod a hb x ≠ a] := + tfae [mem_Ioo_mod p a b, + to_Ico_mod hp a b = to_Ioc_mod hp a b, + to_Ico_mod hp a b + p ≠ to_Ioc_mod hp a b, + to_Ico_mod hp a b ≠ a] := begin tfae_have : 1 → 2, { exact λ ⟨i, hi⟩, - ((to_Ico_mod_eq_iff hb).2 ⟨set.Ioo_subset_Ico_self hi, i, (sub_add_cancel x _).symm⟩).trans - ((to_Ioc_mod_eq_iff hb).2 ⟨set.Ioo_subset_Ioc_self hi, i, (sub_add_cancel x _).symm⟩).symm }, + ((to_Ico_mod_eq_iff hp).2 ⟨set.Ioo_subset_Ico_self hi, i, (sub_add_cancel b _).symm⟩).trans + ((to_Ioc_mod_eq_iff hp).2 ⟨set.Ioo_subset_Ioc_self hi, i, (sub_add_cancel b _).symm⟩).symm }, tfae_have : 2 → 3, - { intro h, rw [h, ne, add_right_eq_self], exact hb.ne' }, + { intro h, rw [h, ne, add_right_eq_self], exact hp.ne' }, tfae_have : 3 → 4, { refine mt (λ h, _), rw [h, eq_comm, to_Ioc_mod_eq_iff, set.right_mem_Ioc], - refine ⟨lt_add_of_pos_right a hb, to_Ico_div a hb x - 1, _⟩, + refine ⟨lt_add_of_pos_right a hp, to_Ico_div hp a b - 1, _⟩, rw [sub_one_zsmul, add_add_add_comm, add_right_neg, add_zero], - conv_lhs { rw [← to_Ico_mod_add_to_Ico_div_zsmul a hb x, h] } }, + conv_lhs { rw [← to_Ico_mod_add_to_Ico_div_zsmul hp a b, h] } }, tfae_have : 4 → 1, - { have h' := to_Ico_mod_mem_Ico a hb x, exact λ h, ⟨_, h'.1.lt_of_ne' h, h'.2⟩ }, + { have h' := to_Ico_mod_mem_Ico hp a b, exact λ h, ⟨_, h'.1.lt_of_ne' h, h'.2⟩ }, tfae_finish, end -variables {a x} +variables {a b} lemma mem_Ioo_mod_iff_to_Ico_mod_eq_to_Ioc_mod : - mem_Ioo_mod a b x ↔ to_Ico_mod a hb x = to_Ioc_mod a hb x := (tfae_mem_Ioo_mod a hb x).out 0 1 + mem_Ioo_mod p a b ↔ to_Ico_mod hp a b = to_Ioc_mod hp a b := (tfae_mem_Ioo_mod hp a b).out 0 1 lemma mem_Ioo_mod_iff_to_Ico_mod_add_period_ne_to_Ioc_mod : - mem_Ioo_mod a b x ↔ to_Ico_mod a hb x + b ≠ to_Ioc_mod a hb x := (tfae_mem_Ioo_mod a hb x).out 0 2 + mem_Ioo_mod p a b ↔ to_Ico_mod hp a b + p ≠ to_Ioc_mod hp a b := (tfae_mem_Ioo_mod hp a b).out 0 2 lemma mem_Ioo_mod_iff_to_Ico_mod_ne_left : - mem_Ioo_mod a b x ↔ to_Ico_mod a hb x ≠ a := (tfae_mem_Ioo_mod a hb x).out 0 3 + mem_Ioo_mod p a b ↔ to_Ico_mod hp a b ≠ a := (tfae_mem_Ioo_mod hp a b).out 0 3 lemma not_mem_Ioo_mod_iff_to_Ico_mod_add_period_eq_to_Ioc_mod : - ¬mem_Ioo_mod a b x ↔ to_Ico_mod a hb x + b = to_Ioc_mod a hb x := -(mem_Ioo_mod_iff_to_Ico_mod_add_period_ne_to_Ioc_mod hb).not_left + ¬mem_Ioo_mod p a b ↔ to_Ico_mod hp a b + p = to_Ioc_mod hp a b := +(mem_Ioo_mod_iff_to_Ico_mod_add_period_ne_to_Ioc_mod hp).not_left -lemma not_mem_Ioo_mod_iff_to_Ico_mod_eq_left : ¬mem_Ioo_mod a b x ↔ to_Ico_mod a hb x = a := -(mem_Ioo_mod_iff_to_Ico_mod_ne_left hb).not_left +lemma not_mem_Ioo_mod_iff_to_Ico_mod_eq_left : ¬mem_Ioo_mod p a b ↔ to_Ico_mod hp a b = a := +(mem_Ioo_mod_iff_to_Ico_mod_ne_left hp).not_left -lemma mem_Ioo_mod_iff_to_Ioc_mod_ne_right : mem_Ioo_mod a b x ↔ to_Ioc_mod a hb x ≠ a + b := +lemma mem_Ioo_mod_iff_to_Ioc_mod_ne_right : mem_Ioo_mod p a b ↔ to_Ioc_mod hp a b ≠ a + p := begin - rw [mem_Ioo_mod_iff_to_Ico_mod_eq_to_Ioc_mod, to_Ico_mod_eq_iff hb], - obtain ⟨h₁, h₂⟩ := to_Ioc_mod_mem_Ioc a hb x, + rw [mem_Ioo_mod_iff_to_Ico_mod_eq_to_Ioc_mod, to_Ico_mod_eq_iff hp], + obtain ⟨h₁, h₂⟩ := to_Ioc_mod_mem_Ioc hp a b, exact ⟨λ h, h.1.2.ne, λ h, ⟨⟨h₁.le, h₂.lt_of_ne h⟩, _, (to_Ioc_mod_add_to_Ioc_div_zsmul _ _ _).symm⟩⟩, end -lemma not_mem_Ioo_mod_iff_to_Ioc_eq_right : ¬mem_Ioo_mod a b x ↔ to_Ioc_mod a hb x = a + b := -(mem_Ioo_mod_iff_to_Ioc_mod_ne_right hb).not_left +lemma not_mem_Ioo_mod_iff_to_Ioc_eq_right : ¬mem_Ioo_mod p a b ↔ to_Ioc_mod hp a b = a + p := +(mem_Ioo_mod_iff_to_Ioc_mod_ne_right hp).not_left lemma mem_Ioo_mod_iff_to_Ico_div_eq_to_Ioc_div : - mem_Ioo_mod a b x ↔ to_Ico_div a hb x = to_Ioc_div a hb x := -by rw [mem_Ioo_mod_iff_to_Ico_mod_eq_to_Ioc_mod hb, - to_Ico_mod, to_Ioc_mod, sub_right_inj, (zsmul_strict_mono_left hb).injective.eq_iff] + mem_Ioo_mod p a b ↔ to_Ico_div hp a b = to_Ioc_div hp a b := +by rw [mem_Ioo_mod_iff_to_Ico_mod_eq_to_Ioc_mod hp, + to_Ico_mod, to_Ioc_mod, sub_right_inj, (zsmul_strict_mono_left hp).injective.eq_iff] lemma mem_Ioo_mod_iff_to_Ico_div_ne_to_Ioc_div_add_one : - mem_Ioo_mod a b x ↔ to_Ico_div a hb x ≠ to_Ioc_div a hb x + 1 := -by rw [mem_Ioo_mod_iff_to_Ico_mod_add_period_ne_to_Ioc_mod hb, ne, ne, to_Ico_mod, to_Ioc_mod, + mem_Ioo_mod p a b ↔ to_Ico_div hp a b ≠ to_Ioc_div hp a b + 1 := +by rw [mem_Ioo_mod_iff_to_Ico_mod_add_period_ne_to_Ioc_mod hp, ne, ne, to_Ico_mod, to_Ioc_mod, ← eq_sub_iff_add_eq, sub_sub, sub_right_inj, ← add_one_zsmul, - (zsmul_strict_mono_left hb).injective.eq_iff] + (zsmul_strict_mono_left hp).injective.eq_iff] lemma not_mem_Ioo_mod_iff_to_Ico_div_eq_to_Ioc_div_add_one : - ¬mem_Ioo_mod a b x ↔ to_Ico_div a hb x = to_Ioc_div a hb x + 1 := -(mem_Ioo_mod_iff_to_Ico_div_ne_to_Ioc_div_add_one hb).not_left + ¬mem_Ioo_mod p a b ↔ to_Ico_div hp a b = to_Ioc_div hp a b + 1 := +(mem_Ioo_mod_iff_to_Ico_div_ne_to_Ioc_div_add_one hp).not_left -include hb +include hp -lemma mem_Ioo_mod_iff_ne_add_zsmul : mem_Ioo_mod a b x ↔ ∀ z : ℤ, x ≠ a + z • b := +lemma mem_Ioo_mod_iff_ne_add_zsmul : mem_Ioo_mod p a b ↔ ∀ z : ℤ, b ≠ a + z • p := begin - rw [mem_Ioo_mod_iff_to_Ico_mod_ne_left hb, ← not_iff_not], + rw [mem_Ioo_mod_iff_to_Ico_mod_ne_left hp, ← not_iff_not], push_neg, split; intro h, { rw ← h, exact ⟨_, (to_Ico_mod_add_to_Ico_div_zsmul _ _ _).symm⟩ }, { rw [to_Ico_mod_eq_iff, set.left_mem_Ico], - exact ⟨lt_add_of_pos_right a hb, h⟩, }, + exact ⟨lt_add_of_pos_right a hp, h⟩, }, end -lemma not_mem_Ioo_mod_iff_eq_add_zsmul : ¬mem_Ioo_mod a b x ↔ ∃ z : ℤ, x = a + z • b := -by simpa only [not_forall, not_ne_iff] using (mem_Ioo_mod_iff_ne_add_zsmul hb).not +lemma not_mem_Ioo_mod_iff_eq_add_zsmul : ¬mem_Ioo_mod p a b ↔ ∃ z : ℤ, b = a + z • p := +by simpa only [not_forall, not_ne_iff] using (mem_Ioo_mod_iff_ne_add_zsmul hp).not lemma not_mem_Ioo_mod_iff_eq_mod_zmultiples : - ¬mem_Ioo_mod a b x ↔ (x : α ⧸ add_subgroup.zmultiples b) = a := -by simp_rw [not_mem_Ioo_mod_iff_eq_add_zsmul hb, quotient_add_group.eq_iff_sub_mem, + ¬mem_Ioo_mod p a b ↔ (b : α ⧸ add_subgroup.zmultiples p) = a := +by simp_rw [not_mem_Ioo_mod_iff_eq_add_zsmul hp, quotient_add_group.eq_iff_sub_mem, add_subgroup.mem_zmultiples_iff, eq_sub_iff_add_eq', eq_comm] lemma mem_Ioo_mod_iff_ne_mod_zmultiples : - mem_Ioo_mod a b x ↔ (x : α ⧸ add_subgroup.zmultiples b) ≠ a := -(not_mem_Ioo_mod_iff_eq_mod_zmultiples hb).not_right + mem_Ioo_mod p a b ↔ (b : α ⧸ add_subgroup.zmultiples p) ≠ a := +(not_mem_Ioo_mod_iff_eq_mod_zmultiples hp).not_right lemma Ico_eq_locus_Ioc_eq_Union_Ioo : - {x | to_Ico_mod a hb x = to_Ioc_mod a hb x} = ⋃ z : ℤ, set.Ioo (a + z • b) (a + b + z • b) := + {b | to_Ico_mod hp a b = to_Ioc_mod hp a b} = ⋃ z : ℤ, set.Ioo (a + z • p) (a + p + z • p) := begin ext1, simp_rw [set.mem_set_of, set.mem_Union, ← set.sub_mem_Ioo_iff_left], - exact (mem_Ioo_mod_iff_to_Ico_mod_eq_to_Ioc_mod hb).symm, + exact (mem_Ioo_mod_iff_to_Ico_mod_eq_to_Ioc_mod hp).symm, end -lemma to_Ioc_div_wcovby_to_Ico_div (a : α) {b : α} (hb : 0 < b) (x : α) : - to_Ioc_div a hb x ⩿ to_Ico_div a hb x := +lemma to_Ioc_div_wcovby_to_Ico_div (a b : α) : to_Ioc_div hp a b ⩿ to_Ico_div hp a b := begin - suffices : to_Ioc_div a hb x = to_Ico_div a hb x ∨ to_Ioc_div a hb x + 1 = to_Ico_div a hb x, + suffices : to_Ioc_div hp a b = to_Ico_div hp a b ∨ to_Ioc_div hp a b + 1 = to_Ico_div hp a b, { rwa [wcovby_iff_eq_or_covby, ←order.succ_eq_iff_covby] }, rw [eq_comm, ←mem_Ioo_mod_iff_to_Ico_div_eq_to_Ioc_div, eq_comm, ←not_mem_Ioo_mod_iff_to_Ico_div_eq_to_Ioc_div_add_one], exact em _, end -lemma to_Ico_mod_le_to_Ioc_mod (a : α) {b : α} (hb : 0 < b) (x : α) : - to_Ico_mod a hb x ≤ to_Ioc_mod a hb x := +lemma to_Ico_mod_le_to_Ioc_mod (a b : α) : to_Ico_mod hp a b ≤ to_Ioc_mod hp a b := begin rw [to_Ico_mod, to_Ioc_mod, sub_le_sub_iff_left], - exact zsmul_mono_left hb.le (to_Ioc_div_wcovby_to_Ico_div _ _ _).le + exact zsmul_mono_left hp.le (to_Ioc_div_wcovby_to_Ico_div _ _ _).le end -lemma to_Ioc_mod_le_to_Ico_mod_add (a : α) {b : α} (hb : 0 < b) (x : α) : - to_Ioc_mod a hb x ≤ to_Ico_mod a hb x + b := +lemma to_Ioc_mod_le_to_Ico_mod_add (a b : α) : to_Ioc_mod hp a b ≤ to_Ico_mod hp a b + p := begin rw [to_Ico_mod, to_Ioc_mod, sub_add, sub_le_sub_iff_left, sub_le_iff_le_add, ←add_one_zsmul, - (zsmul_strict_mono_left hb).le_iff_le], + (zsmul_strict_mono_left hp).le_iff_le], apply (to_Ioc_div_wcovby_to_Ico_div _ _ _).le_succ, end end Ico_Ioc -lemma to_Ico_mod_eq_self {a b x : α} (hb : 0 < b) : to_Ico_mod a hb x = x ↔ x ∈ set.Ico a (a + b) := -begin - rw [to_Ico_mod_eq_iff, and_iff_left], - refine ⟨0, _⟩, - simp -end +lemma to_Ico_mod_eq_self : to_Ico_mod hp a b = b ↔ b ∈ set.Ico a (a + p) := +by { rw [to_Ico_mod_eq_iff, and_iff_left], exact ⟨0, by simp⟩ } -lemma to_Ioc_mod_eq_self {a b x : α} (hb : 0 < b) : to_Ioc_mod a hb x = x ↔ x ∈ set.Ioc a (a + b) := -begin - rw [to_Ioc_mod_eq_iff, and_iff_left], - refine ⟨0, _⟩, - simp -end +lemma to_Ioc_mod_eq_self : to_Ioc_mod hp a b = b ↔ b ∈ set.Ioc a (a + p) := +by { rw [to_Ioc_mod_eq_iff, and_iff_left], exact ⟨0, by simp⟩ } -@[simp] lemma to_Ico_mod_to_Ico_mod (a₁ a₂ : α) {b : α} (hb : 0 < b) (x : α) : - to_Ico_mod a₁ hb (to_Ico_mod a₂ hb x) = to_Ico_mod a₁ hb x := -begin - rw to_Ico_mod_eq_to_Ico_mod, - exact ⟨to_Ico_div a₂ hb x, self_sub_to_Ico_mod a₂ hb x⟩ -end +@[simp] lemma to_Ico_mod_to_Ico_mod (a₁ a₂ b : α) : + to_Ico_mod hp a₁ (to_Ico_mod hp a₂ b) = to_Ico_mod hp a₁ b := +(to_Ico_mod_eq_to_Ico_mod _).2 ⟨to_Ico_div hp a₂ b, self_sub_to_Ico_mod hp a₂ b⟩ -@[simp] lemma to_Ico_mod_to_Ioc_mod (a₁ a₂ : α) {b : α} (hb : 0 < b) (x : α) : - to_Ico_mod a₁ hb (to_Ioc_mod a₂ hb x) = to_Ico_mod a₁ hb x := -begin - rw to_Ico_mod_eq_to_Ico_mod, - exact ⟨to_Ioc_div a₂ hb x, self_sub_to_Ioc_mod a₂ hb x⟩ -end +@[simp] lemma to_Ico_mod_to_Ioc_mod (a₁ a₂ b : α) : + to_Ico_mod hp a₁ (to_Ioc_mod hp a₂ b) = to_Ico_mod hp a₁ b := +(to_Ico_mod_eq_to_Ico_mod _).2 ⟨to_Ioc_div hp a₂ b, self_sub_to_Ioc_mod hp a₂ b⟩ -@[simp] lemma to_Ioc_mod_to_Ioc_mod (a₁ a₂ : α) {b : α} (hb : 0 < b) (x : α) : - to_Ioc_mod a₁ hb (to_Ioc_mod a₂ hb x) = to_Ioc_mod a₁ hb x := -begin - rw to_Ioc_mod_eq_to_Ioc_mod, - exact ⟨to_Ioc_div a₂ hb x, self_sub_to_Ioc_mod a₂ hb x⟩ -end +@[simp] lemma to_Ioc_mod_to_Ioc_mod (a₁ a₂ b : α) : + to_Ioc_mod hp a₁ (to_Ioc_mod hp a₂ b) = to_Ioc_mod hp a₁ b := +(to_Ioc_mod_eq_to_Ioc_mod _).2 ⟨to_Ioc_div hp a₂ b, self_sub_to_Ioc_mod hp a₂ b⟩ -@[simp] lemma to_Ioc_mod_to_Ico_mod (a₁ a₂ : α) {b : α} (hb : 0 < b) (x : α) : - to_Ioc_mod a₁ hb (to_Ico_mod a₂ hb x) = to_Ioc_mod a₁ hb x := -begin - rw to_Ioc_mod_eq_to_Ioc_mod, - exact ⟨to_Ico_div a₂ hb x, self_sub_to_Ico_mod a₂ hb x⟩ -end +@[simp] lemma to_Ioc_mod_to_Ico_mod (a₁ a₂ b : α) : + to_Ioc_mod hp a₁ (to_Ico_mod hp a₂ b) = to_Ioc_mod hp a₁ b := +(to_Ioc_mod_eq_to_Ioc_mod _).2 ⟨to_Ico_div hp a₂ b, self_sub_to_Ico_mod hp a₂ b⟩ -lemma to_Ico_mod_periodic (a : α) {b : α} (hb : 0 < b) : function.periodic (to_Ico_mod a hb) b := -to_Ico_mod_add_right a hb +lemma to_Ico_mod_periodic (a : α) : function.periodic (to_Ico_mod hp a) p := +to_Ico_mod_add_right hp a -lemma to_Ioc_mod_periodic (a : α) {b : α} (hb : 0 < b) : function.periodic (to_Ioc_mod a hb) b := -to_Ioc_mod_add_right a hb +lemma to_Ioc_mod_periodic (a : α) : function.periodic (to_Ioc_mod hp a) p := +to_Ioc_mod_add_right hp a /-- `to_Ico_mod` as an equiv from the quotient. -/ @[simps symm_apply] -def quotient_add_group.equiv_Ico_mod (a : α) {b : α} (hb : 0 < b) : - (α ⧸ add_subgroup.zmultiples b) ≃ set.Ico a (a + b) := -{ to_fun := λ x, ⟨(to_Ico_mod_periodic a hb).lift x, - quotient_add_group.induction_on' x $ to_Ico_mod_mem_Ico a hb⟩, +def quotient_add_group.equiv_Ico_mod (a : α) : + (α ⧸ add_subgroup.zmultiples p) ≃ set.Ico a (a + p) := +{ to_fun := λ b, ⟨(to_Ico_mod_periodic hp a).lift b, + quotient_add_group.induction_on' b $ to_Ico_mod_mem_Ico hp a⟩, inv_fun := coe, - right_inv := λ x, subtype.ext $ (to_Ico_mod_eq_self hb).mpr x.prop, - left_inv := λ x, begin - induction x using quotient_add_group.induction_on', + right_inv := λ b, subtype.ext $ (to_Ico_mod_eq_self hp).mpr b.prop, + left_inv := λ b, begin + induction b using quotient_add_group.induction_on', dsimp, rw [quotient_add_group.eq_iff_sub_mem, to_Ico_mod_sub_self], apply add_subgroup.zsmul_mem_zmultiples, end } @[simp] -lemma quotient_add_group.equiv_Ico_mod_coe (a : α) {b : α} (hb : 0 < b) (x : α) : - quotient_add_group.equiv_Ico_mod a hb ↑x = ⟨to_Ico_mod a hb x, to_Ico_mod_mem_Ico a hb _⟩ := +lemma quotient_add_group.equiv_Ico_mod_coe (a b : α) : + quotient_add_group.equiv_Ico_mod hp a ↑b = ⟨to_Ico_mod hp a b, to_Ico_mod_mem_Ico hp a _⟩ := rfl /-- `to_Ioc_mod` as an equiv from the quotient. -/ @[simps symm_apply] -def quotient_add_group.equiv_Ioc_mod (a : α) {b : α} (hb : 0 < b) : - (α ⧸ add_subgroup.zmultiples b) ≃ set.Ioc a (a + b) := -{ to_fun := λ x, ⟨(to_Ioc_mod_periodic a hb).lift x, - quotient_add_group.induction_on' x $ to_Ioc_mod_mem_Ioc a hb⟩, +def quotient_add_group.equiv_Ioc_mod (a : α) : + (α ⧸ add_subgroup.zmultiples p) ≃ set.Ioc a (a + p) := +{ to_fun := λ b, ⟨(to_Ioc_mod_periodic hp a).lift b, + quotient_add_group.induction_on' b $ to_Ioc_mod_mem_Ioc hp a⟩, inv_fun := coe, - right_inv := λ x, subtype.ext $ (to_Ioc_mod_eq_self hb).mpr x.prop, - left_inv := λ x, begin - induction x using quotient_add_group.induction_on', + right_inv := λ b, subtype.ext $ (to_Ioc_mod_eq_self hp).mpr b.prop, + left_inv := λ b, begin + induction b using quotient_add_group.induction_on', dsimp, rw [quotient_add_group.eq_iff_sub_mem, to_Ioc_mod_sub_self], apply add_subgroup.zsmul_mem_zmultiples, end } @[simp] -lemma quotient_add_group.equiv_Ioc_mod_coe (a : α) {b : α} (hb : 0 < b) (x : α) : - quotient_add_group.equiv_Ioc_mod a hb ↑x = ⟨to_Ioc_mod a hb x, to_Ioc_mod_mem_Ioc a hb _⟩ := +lemma quotient_add_group.equiv_Ioc_mod_coe (a b : α) : + quotient_add_group.equiv_Ioc_mod hp a ↑b = ⟨to_Ioc_mod hp a b, to_Ioc_mod_mem_Ioc hp a _⟩ := rfl end linear_ordered_add_comm_group section linear_ordered_field -variables {α : Type*} [linear_ordered_field α] [floor_ring α] +variables {α : Type*} [linear_ordered_field α] [floor_ring α] {p : α} (hp : 0 < p) -lemma to_Ico_div_eq_floor (a : α) {b : α} (hb : 0 < b) (x : α) : - to_Ico_div a hb x = ⌊(x - a) / b⌋ := +lemma to_Ico_div_eq_floor (a b : α) : to_Ico_div hp a b = ⌊(b - a) / p⌋ := begin - refine (eq_to_Ico_div_of_sub_zsmul_mem_Ico hb _).symm, + refine to_Ico_div_eq_of_sub_zsmul_mem_Ico hp _, rw [set.mem_Ico, zsmul_eq_mul, ←sub_nonneg, add_comm, sub_right_comm, ←sub_lt_iff_lt_add, sub_right_comm _ _ a], - exact ⟨int.sub_floor_div_mul_nonneg _ hb, int.sub_floor_div_mul_lt _ hb⟩, + exact ⟨int.sub_floor_div_mul_nonneg _ hp, int.sub_floor_div_mul_lt _ hp⟩, end -lemma to_Ioc_div_eq_neg_floor (a : α) {b : α} (hb : 0 < b) (x : α) : - to_Ioc_div a hb x = -⌊(a + b - x) / b⌋ := +lemma to_Ioc_div_eq_neg_floor (a b : α) : to_Ioc_div hp a b = -⌊(a + p - b) / p⌋ := begin - refine (eq_to_Ioc_div_of_sub_zsmul_mem_Ioc hb _).symm, + refine to_Ioc_div_eq_of_sub_zsmul_mem_Ioc hp _, rw [set.mem_Ioc, zsmul_eq_mul, int.cast_neg, neg_mul, sub_neg_eq_add, ←sub_nonneg, sub_add_eq_sub_sub], - refine ⟨_, int.sub_floor_div_mul_nonneg _ hb⟩, - rw [←add_lt_add_iff_right b, add_assoc, add_comm x, ←sub_lt_iff_lt_add, add_comm (_ * _), + refine ⟨_, int.sub_floor_div_mul_nonneg _ hp⟩, + rw [←add_lt_add_iff_right p, add_assoc, add_comm b, ←sub_lt_iff_lt_add, add_comm (_ * _), ←sub_lt_iff_lt_add], - exact int.sub_floor_div_mul_lt _ hb + exact int.sub_floor_div_mul_lt _ hp end -lemma to_Ico_div_zero_one (x : α) : to_Ico_div (0 : α) zero_lt_one x = ⌊x⌋ := +lemma to_Ico_div_zero_one (b : α) : to_Ico_div (zero_lt_one' α) 0 b = ⌊b⌋ := by simp [to_Ico_div_eq_floor] -lemma to_Ico_mod_eq_add_fract_mul (a : α) {b : α} (hb : 0 < b) (x : α) : - to_Ico_mod a hb x = a + int.fract ((x - a) / b) * b := +lemma to_Ico_mod_eq_add_fract_mul (a b : α) : to_Ico_mod hp a b = a + int.fract ((b - a) / p) * p := begin rw [to_Ico_mod, to_Ico_div_eq_floor, int.fract], - field_simp [hb.ne.symm], + field_simp [hp.ne.symm], ring end -lemma to_Ico_mod_eq_fract_mul {b : α} (hb : 0 < b) (x : α) : - to_Ico_mod 0 hb x = int.fract (x / b) * b := +lemma to_Ico_mod_eq_fract_mul (b : α) : to_Ico_mod hp 0 b = int.fract (b / p) * p := by simp [to_Ico_mod_eq_add_fract_mul] -lemma to_Ioc_mod_eq_sub_fract_mul (a : α) {b : α} (hb : 0 < b) (x : α) : - to_Ioc_mod a hb x = a + b - int.fract ((a + b - x) / b) * b := +lemma to_Ioc_mod_eq_sub_fract_mul (a b : α) : + to_Ioc_mod hp a b = a + p - int.fract ((a + p - b) / p) * p := begin rw [to_Ioc_mod, to_Ioc_div_eq_neg_floor, int.fract], - field_simp [hb.ne.symm], + field_simp [hp.ne.symm], ring end -lemma to_Ico_mod_zero_one (x : α) : to_Ico_mod (0 : α) zero_lt_one x = int.fract x := +lemma to_Ico_mod_zero_one (b : α) : to_Ico_mod (zero_lt_one' α) 0 b = int.fract b := by simp [to_Ico_mod_eq_add_fract_mul] end linear_ordered_field @@ -744,39 +593,39 @@ open set int section linear_ordered_add_comm_group -variables {α : Type*} [linear_ordered_add_comm_group α] [archimedean α] (a : α) {b : α} (hb : 0 < b) -include hb +variables {α : Type*} [linear_ordered_add_comm_group α] [archimedean α] {p : α} (hp : 0 < p) (a : α) +include hp -lemma Union_Ioc_add_zsmul : (⋃ (n : ℤ), Ioc (a + n • b) (a + (n + 1) • b)) = univ := +lemma Union_Ioc_add_zsmul : (⋃ (n : ℤ), Ioc (a + n • p) (a + (n + 1) • p)) = univ := begin - refine eq_univ_iff_forall.mpr (λ x, mem_Union.mpr _), - rcases sub_to_Ioc_div_zsmul_mem_Ioc a hb x with ⟨hl, hr⟩, - refine ⟨to_Ioc_div a hb x, ⟨lt_sub_iff_add_lt.mp hl, _⟩⟩, + refine eq_univ_iff_forall.mpr (λ b, mem_Union.mpr _), + rcases sub_to_Ioc_div_zsmul_mem_Ioc hp a b with ⟨hl, hr⟩, + refine ⟨to_Ioc_div hp a b, ⟨lt_sub_iff_add_lt.mp hl, _⟩⟩, rw [add_smul, one_smul, ←add_assoc], convert sub_le_iff_le_add.mp hr using 1, abel, end -lemma Union_Ico_add_zsmul : (⋃ (n : ℤ), Ico (a + n • b) (a + (n + 1) • b)) = univ := +lemma Union_Ico_add_zsmul : (⋃ (n : ℤ), Ico (a + n • p) (a + (n + 1) • p)) = univ := begin - refine eq_univ_iff_forall.mpr (λ x, mem_Union.mpr _), - rcases sub_to_Ico_div_zsmul_mem_Ico a hb x with ⟨hl, hr⟩, - refine ⟨to_Ico_div a hb x, ⟨le_sub_iff_add_le.mp hl, _⟩⟩, + refine eq_univ_iff_forall.mpr (λ b, mem_Union.mpr _), + rcases sub_to_Ico_div_zsmul_mem_Ico hp a b with ⟨hl, hr⟩, + refine ⟨to_Ico_div hp a b, ⟨le_sub_iff_add_le.mp hl, _⟩⟩, rw [add_smul, one_smul, ←add_assoc], convert sub_lt_iff_lt_add.mp hr using 1, abel, end -lemma Union_Icc_add_zsmul : (⋃ (n : ℤ), Icc (a + n • b) (a + (n + 1) • b)) = univ := -by simpa only [Union_Ioc_add_zsmul a hb, univ_subset_iff] using - Union_mono (λ n : ℤ, (Ioc_subset_Icc_self : Ioc (a + n • b) (a + (n + 1) • b) ⊆ Icc _ _)) +lemma Union_Icc_add_zsmul : (⋃ (n : ℤ), Icc (a + n • p) (a + (n + 1) • p)) = univ := +by simpa only [Union_Ioc_add_zsmul hp a, univ_subset_iff] using + Union_mono (λ n : ℤ, (Ioc_subset_Icc_self : Ioc (a + n • p) (a + (n + 1) • p) ⊆ Icc _ _)) -lemma Union_Ioc_zsmul : (⋃ (n : ℤ), Ioc (n • b) ((n + 1) • b)) = univ := -by simpa only [zero_add] using Union_Ioc_add_zsmul 0 hb +lemma Union_Ioc_zsmul : (⋃ (n : ℤ), Ioc (n • p) ((n + 1) • p)) = univ := +by simpa only [zero_add] using Union_Ioc_add_zsmul hp 0 -lemma Union_Ico_zsmul : (⋃ (n : ℤ), Ico (n • b) ((n + 1) • b)) = univ := -by simpa only [zero_add] using Union_Ico_add_zsmul 0 hb +lemma Union_Ico_zsmul : (⋃ (n : ℤ), Ico (n • p) ((n + 1) • p)) = univ := +by simpa only [zero_add] using Union_Ico_add_zsmul hp 0 -lemma Union_Icc_zsmul : (⋃ (n : ℤ), Icc (n • b) ((n + 1) • b)) = univ := -by simpa only [zero_add] using Union_Icc_add_zsmul 0 hb +lemma Union_Icc_zsmul : (⋃ (n : ℤ), Icc (n • p) ((n + 1) • p)) = univ := +by simpa only [zero_add] using Union_Icc_add_zsmul hp 0 end linear_ordered_add_comm_group @@ -786,15 +635,15 @@ variables {α : Type*} [linear_ordered_ring α] [archimedean α] (a : α) lemma Union_Ioc_add_int_cast : (⋃ (n : ℤ), Ioc (a + n) (a + n + 1)) = set.univ := by simpa only [zsmul_one, int.cast_add, int.cast_one, ←add_assoc] - using Union_Ioc_add_zsmul a zero_lt_one + using Union_Ioc_add_zsmul zero_lt_one a lemma Union_Ico_add_int_cast : (⋃ (n : ℤ), Ico (a + n) (a + n + 1)) = set.univ := by simpa only [zsmul_one, int.cast_add, int.cast_one, ←add_assoc] - using Union_Ico_add_zsmul a zero_lt_one + using Union_Ico_add_zsmul zero_lt_one a lemma Union_Icc_add_int_cast : (⋃ (n : ℤ), Icc (a + n) (a + n + 1)) = set.univ := by simpa only [zsmul_one, int.cast_add, int.cast_one, ←add_assoc] - using Union_Icc_add_zsmul a (zero_lt_one' α) + using Union_Icc_add_zsmul zero_lt_one a variables (α) diff --git a/src/analysis/special_functions/complex/arg.lean b/src/analysis/special_functions/complex/arg.lean index b2d28ee6d520c..76fa41d560635 100644 --- a/src/analysis/special_functions/complex/arg.lean +++ b/src/analysis/special_functions/complex/arg.lean @@ -389,17 +389,17 @@ begin end lemma arg_mul_cos_add_sin_mul_I_eq_to_Ioc_mod {r : ℝ} (hr : 0 < r) (θ : ℝ) : - arg (r * (cos θ + sin θ * I)) = to_Ioc_mod (-π) real.two_pi_pos θ := + arg (r * (cos θ + sin θ * I)) = to_Ioc_mod real.two_pi_pos (-π) θ := begin - have hi : to_Ioc_mod (-π) real.two_pi_pos θ ∈ Ioc (-π) π, - { convert to_Ioc_mod_mem_Ioc _ real.two_pi_pos _, + have hi : to_Ioc_mod real.two_pi_pos (-π) θ ∈ Ioc (-π) π, + { convert to_Ioc_mod_mem_Ioc _ _ _, ring }, convert arg_mul_cos_add_sin_mul_I hr hi using 3, simp [to_Ioc_mod, cos_sub_int_mul_two_pi, sin_sub_int_mul_two_pi] end lemma arg_cos_add_sin_mul_I_eq_to_Ioc_mod (θ : ℝ) : - arg (cos θ + sin θ * I) = to_Ioc_mod (-π) real.two_pi_pos θ := + arg (cos θ + sin θ * I) = to_Ioc_mod real.two_pi_pos (-π) θ := by rw [←one_mul (_ + _), ←of_real_one, arg_mul_cos_add_sin_mul_I_eq_to_Ioc_mod zero_lt_one] lemma arg_mul_cos_add_sin_mul_I_sub {r : ℝ} (hr : 0 < r) (θ : ℝ) : diff --git a/src/analysis/special_functions/trigonometric/angle.lean b/src/analysis/special_functions/trigonometric/angle.lean index 95c8108478a21..f7eed78ad9d96 100644 --- a/src/analysis/special_functions/trigonometric/angle.lean +++ b/src/analysis/special_functions/trigonometric/angle.lean @@ -404,26 +404,25 @@ begin exact abs_cos_eq_of_two_nsmul_eq h end -@[simp] lemma coe_to_Ico_mod (θ ψ : ℝ) : ↑(to_Ico_mod ψ two_pi_pos θ) = (θ : angle) := +@[simp] lemma coe_to_Ico_mod (θ ψ : ℝ) : ↑(to_Ico_mod two_pi_pos ψ θ) = (θ : angle) := begin rw angle_eq_iff_two_pi_dvd_sub, - refine ⟨-to_Ico_div ψ two_pi_pos θ, _⟩, + refine ⟨-to_Ico_div two_pi_pos ψ θ, _⟩, rw [to_Ico_mod_sub_self, zsmul_eq_mul, mul_comm] end -@[simp] lemma coe_to_Ioc_mod (θ ψ : ℝ) : ↑(to_Ioc_mod ψ two_pi_pos θ) = (θ : angle) := +@[simp] lemma coe_to_Ioc_mod (θ ψ : ℝ) : ↑(to_Ioc_mod two_pi_pos ψ θ) = (θ : angle) := begin rw angle_eq_iff_two_pi_dvd_sub, - refine ⟨-to_Ioc_div ψ two_pi_pos θ, _⟩, + refine ⟨-to_Ioc_div two_pi_pos ψ θ, _⟩, rw [to_Ioc_mod_sub_self, zsmul_eq_mul, mul_comm] end /-- Convert a `real.angle` to a real number in the interval `Ioc (-π) π`. -/ def to_real (θ : angle) : ℝ := -(to_Ioc_mod_periodic (-π) two_pi_pos).lift θ +(to_Ioc_mod_periodic two_pi_pos (-π)).lift θ -lemma to_real_coe (θ : ℝ) : (θ : angle).to_real = to_Ioc_mod (-π) two_pi_pos θ := -rfl +lemma to_real_coe (θ : ℝ) : (θ : angle).to_real = to_Ioc_mod two_pi_pos (-π) θ := rfl lemma to_real_coe_eq_self_iff {θ : ℝ} : (θ : angle).to_real = θ ↔ -π < θ ∧ θ ≤ π := begin @@ -455,13 +454,13 @@ end lemma neg_pi_lt_to_real (θ : angle) : -π < θ.to_real := begin induction θ using real.angle.induction_on, - exact left_lt_to_Ioc_mod _ two_pi_pos _ + exact left_lt_to_Ioc_mod _ _ _ end lemma to_real_le_pi (θ : angle) : θ.to_real ≤ π := begin induction θ using real.angle.induction_on, - convert to_Ioc_mod_le_right _ two_pi_pos _, + convert to_Ioc_mod_le_right two_pi_pos _ _, ring end @@ -471,7 +470,7 @@ abs_le.2 ⟨(neg_pi_lt_to_real _).le, to_real_le_pi _⟩ lemma to_real_mem_Ioc (θ : angle) : θ.to_real ∈ set.Ioc (-π) π := ⟨neg_pi_lt_to_real _, to_real_le_pi _⟩ -@[simp] lemma to_Ioc_mod_to_real (θ : angle): to_Ioc_mod (-π) two_pi_pos θ.to_real = θ.to_real := +@[simp] lemma to_Ioc_mod_to_real (θ : angle): to_Ioc_mod two_pi_pos (-π) θ.to_real = θ.to_real := begin induction θ using real.angle.induction_on, rw to_real_coe, diff --git a/src/topology/instances/add_circle.lean b/src/topology/instances/add_circle.lean index 4e626f5d7c06e..40a4bf35aa9d7 100644 --- a/src/topology/instances/add_circle.lean +++ b/src/topology/instances/add_circle.lean @@ -57,17 +57,17 @@ variables {𝕜 B : Type*} section continuity variables [linear_ordered_add_comm_group 𝕜] [archimedean 𝕜] - [topological_space 𝕜] [order_topology 𝕜] (a : 𝕜) {p : 𝕜} (hp : 0 < p) (x : 𝕜) + [topological_space 𝕜] [order_topology 𝕜] {p : 𝕜} (hp : 0 < p) (a x : 𝕜) -lemma continuous_right_to_Ico_mod : continuous_within_at (to_Ico_mod a hp) (Ici x) x := +lemma continuous_right_to_Ico_mod : continuous_within_at (to_Ico_mod hp a) (Ici x) x := begin intros s h, rw [filter.mem_map, mem_nhds_within_iff_exists_mem_nhds_inter], haveI : nontrivial 𝕜 := ⟨⟨0, p, hp.ne⟩⟩, simp_rw mem_nhds_iff_exists_Ioo_subset at h ⊢, obtain ⟨l, u, hxI, hIs⟩ := h, - let d := to_Ico_div a hp x • p, - have hd := to_Ico_mod_mem_Ico a hp x, + let d := to_Ico_div hp a x • p, + have hd := to_Ico_mod_mem_Ico hp a x, simp_rw [subset_def, mem_inter_iff], refine ⟨_, ⟨l + d, min (a + p) u + d, _, λ x, id⟩, λ y, _⟩; simp_rw [← sub_mem_Ioo_iff_left, mem_Ioo, lt_min_iff], @@ -77,28 +77,28 @@ begin exacts [⟨h.1, h.2.2⟩, ⟨hd.1.trans (sub_le_sub_right h' _), h.2.1⟩] }, end -lemma continuous_left_to_Ioc_mod : continuous_within_at (to_Ioc_mod a hp) (Iic x) x := +lemma continuous_left_to_Ioc_mod : continuous_within_at (to_Ioc_mod hp a) (Iic x) x := begin rw (funext (λ y, eq.trans (by rw neg_neg) $ to_Ioc_mod_neg _ _ _) : - to_Ioc_mod a hp = (λ x, p - x) ∘ to_Ico_mod (-a) hp ∘ has_neg.neg), + to_Ioc_mod hp a = (λ x, p - x) ∘ to_Ico_mod hp (-a) ∘ has_neg.neg), exact ((continuous_sub_left _).continuous_at.comp_continuous_within_at $ (continuous_right_to_Ico_mod _ _ _).comp continuous_neg.continuous_within_at $ λ y, neg_le_neg), end variables {x} (hx : (x : 𝕜 ⧸ zmultiples p) ≠ a) -lemma to_Ico_mod_eventually_eq_to_Ioc_mod : to_Ico_mod a hp =ᶠ[𝓝 x] to_Ioc_mod a hp := +lemma to_Ico_mod_eventually_eq_to_Ioc_mod : to_Ico_mod hp a =ᶠ[𝓝 x] to_Ioc_mod hp a := is_open.mem_nhds (by {rw Ico_eq_locus_Ioc_eq_Union_Ioo, exact is_open_Union (λ i, is_open_Ioo)}) $ (mem_Ioo_mod_iff_to_Ico_mod_eq_to_Ioc_mod hp).1 ((mem_Ioo_mod_iff_ne_mod_zmultiples hp).2 hx) -lemma continuous_at_to_Ico_mod : continuous_at (to_Ico_mod a hp) x := -let h := to_Ico_mod_eventually_eq_to_Ioc_mod a hp hx in continuous_at_iff_continuous_left_right.2 $ - ⟨(continuous_left_to_Ioc_mod a hp x).congr_of_eventually_eq - (h.filter_mono nhds_within_le_nhds) h.eq_of_nhds, continuous_right_to_Ico_mod a hp x⟩ +lemma continuous_at_to_Ico_mod : continuous_at (to_Ico_mod hp a) x := +let h := to_Ico_mod_eventually_eq_to_Ioc_mod hp a hx in continuous_at_iff_continuous_left_right.2 $ + ⟨(continuous_left_to_Ioc_mod hp a x).congr_of_eventually_eq + (h.filter_mono nhds_within_le_nhds) h.eq_of_nhds, continuous_right_to_Ico_mod hp a x⟩ -lemma continuous_at_to_Ioc_mod : continuous_at (to_Ioc_mod a hp) x := -let h := to_Ico_mod_eventually_eq_to_Ioc_mod a hp hx in continuous_at_iff_continuous_left_right.2 $ - ⟨continuous_left_to_Ioc_mod a hp x, (continuous_right_to_Ico_mod a hp x).congr_of_eventually_eq +lemma continuous_at_to_Ioc_mod : continuous_at (to_Ioc_mod hp a) x := +let h := to_Ico_mod_eventually_eq_to_Ioc_mod hp a hx in continuous_at_iff_continuous_left_right.2 $ + ⟨continuous_left_to_Ioc_mod hp a x, (continuous_right_to_Ico_mod hp a x).congr_of_eventually_eq (h.symm.filter_mono nhds_within_le_nhds) h.symm.eq_of_nhds⟩ end continuity @@ -157,11 +157,11 @@ variables (a : 𝕜) [archimedean 𝕜] /-- The equivalence between `add_circle p` and the half-open interval `[a, a + p)`, whose inverse is the natural quotient map. -/ -def equiv_Ico : add_circle p ≃ Ico a (a + p) := quotient_add_group.equiv_Ico_mod a hp.out +def equiv_Ico : add_circle p ≃ Ico a (a + p) := quotient_add_group.equiv_Ico_mod hp.out a /-- The equivalence between `add_circle p` and the half-open interval `(a, a + p]`, whose inverse is the natural quotient map. -/ -def equiv_Ioc : add_circle p ≃ Ioc a (a + p) := quotient_add_group.equiv_Ioc_mod a hp.out +def equiv_Ioc : add_circle p ≃ Ioc a (a + p) := quotient_add_group.equiv_Ioc_mod hp.out a /-- Given a function on `𝕜`, return the unique function on `add_circle p` agreeing with `f` on `[a, a + p)`. -/ @@ -218,14 +218,14 @@ lemma continuous_at_equiv_Ico : continuous_at (equiv_Ico p a) x := begin induction x using quotient_add_group.induction_on', rw [continuous_at, filter.tendsto, quotient_add_group.nhds_eq, filter.map_map], - apply continuous_at.cod_restrict, exact continuous_at_to_Ico_mod a hp.out hx, + exact (continuous_at_to_Ico_mod hp.out a hx).cod_restrict _, end lemma continuous_at_equiv_Ioc : continuous_at (equiv_Ioc p a) x := begin induction x using quotient_add_group.induction_on', rw [continuous_at, filter.tendsto, quotient_add_group.nhds_eq, filter.map_map], - apply continuous_at.cod_restrict, exact continuous_at_to_Ioc_mod a hp.out hx, + exact (continuous_at_to_Ioc_mod hp.out a hx).cod_restrict _, end end continuity @@ -492,13 +492,13 @@ def equiv_Icc_quot : 𝕋 ≃ quot (endpoint_ident p a) := apply congr_arg subtype.val ((equiv_Ico p a).right_inv ⟨x, hx.1, hx.2.lt_of_ne h⟩) } } lemma equiv_Icc_quot_comp_mk_eq_to_Ico_mod : equiv_Icc_quot p a ∘ quotient.mk' = - λ x, quot.mk _ ⟨to_Ico_mod a hp.out x, Ico_subset_Icc_self $ to_Ico_mod_mem_Ico a _ x⟩ := rfl + λ x, quot.mk _ ⟨to_Ico_mod hp.out a x, Ico_subset_Icc_self $ to_Ico_mod_mem_Ico _ _ x⟩ := rfl lemma equiv_Icc_quot_comp_mk_eq_to_Ioc_mod : equiv_Icc_quot p a ∘ quotient.mk' = - λ x, quot.mk _ ⟨to_Ioc_mod a hp.out x, Ioc_subset_Icc_self $ to_Ioc_mod_mem_Ioc a _ x⟩ := + λ x, quot.mk _ ⟨to_Ioc_mod hp.out a x, Ioc_subset_Icc_self $ to_Ioc_mod_mem_Ioc _ _ x⟩ := begin rw equiv_Icc_quot_comp_mk_eq_to_Ico_mod, funext, - by_cases mem_Ioo_mod a p x, + by_cases mem_Ioo_mod p a x, { simp_rw (mem_Ioo_mod_iff_to_Ico_mod_eq_to_Ioc_mod hp.out).1 h }, { simp_rw [not_imp_comm.1 (mem_Ioo_mod_iff_to_Ico_mod_ne_left hp.out).2 h, not_imp_comm.1 (mem_Ioo_mod_iff_to_Ioc_mod_ne_right hp.out).2 h], From 620ba06c7f173d79b3cad91294babde6b4eab79c Mon Sep 17 00:00:00 2001 From: Apurva Nakade Date: Tue, 2 May 2023 01:48:43 +0000 Subject: [PATCH 0921/1291] feat(analysis/convex/cone/proper): define closure of a convex cone (#16335) Part of #16266 We define the closure of a convex cone. This definition is only used for defining maps between proper cones and hence is in this file and not cone/basic.lean Next todo: define proper cones and add some proper cone API Co-authored-by: Apurva Nakade --- docs/references.bib | 6 ++++ src/analysis/convex/cone/proper.lean | 53 ++++++++++++++++++++++++++++ 2 files changed, 59 insertions(+) create mode 100644 src/analysis/convex/cone/proper.lean diff --git a/docs/references.bib b/docs/references.bib index da809a2690cf7..d57e87b338103 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -956,6 +956,12 @@ @InProceedings{ Gallier2011Notes url = {https://www.cis.upenn.edu/~cis610/diffgeom-n.pdf} } +@Unpublished{ gartnerMatousek, + title = {Cone Programming}, + author = {B. G{\"{a}}rtner and J. Matousek}, + url = {https://ti.inf.ethz.ch/ew/lehre/ApproxSDP09/notes/conelp.pdf} +} + @Article{ ghys87:groupes, author = {Étienne Ghys}, title = {Groupes d'homéomorphismes du cercle et cohomologie diff --git a/src/analysis/convex/cone/proper.lean b/src/analysis/convex/cone/proper.lean new file mode 100644 index 0000000000000..2bc6da0856b11 --- /dev/null +++ b/src/analysis/convex/cone/proper.lean @@ -0,0 +1,53 @@ +/- +Copyright (c) 2022 Apurva Nakade All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Apurva Nakade +-/ +import analysis.inner_product_space.adjoint + +/-! + +We define a proper cone as a nonempty, closed, convex cone. Proper cones are used in defining conic +programs which generalize linear programs. A linear program is a conic program for the positive +cone. We then prove Farkas' lemma for conic programs following the proof in the reference below. +Farkas' lemma is equivalent to strong duality. So, once have the definitions of conic programs and +linear programs, the results from this file can be used to prove duality theorems. + +## TODO + +In the next few PRs (already sorry-free), we will add the definition and prove several properties +of proper cones and finally prove the cone version of Farkas' lemma (2.3.4 in the reference). + +The next steps are: +- Define primal and dual cone programs and prove weak duality. +- Prove regular and strong duality for cone programs using Farkas' lemma (see reference). +- Define linear programs and prove LP duality as a special case of cone duality. + +## References + +- [B. Gartner and J. Matousek, Cone Programming][gartnerMatousek] + +-/ + +open continuous_linear_map filter + +namespace convex_cone + +variables {E : Type*} [add_comm_monoid E] [has_smul ℝ E] [topological_space E] + [has_continuous_const_smul ℝ E] [has_continuous_add E] + +/-- The closure of a convex cone inside a real inner product space is a convex cone. This +construction is mainly used for defining maps between proper cones. -/ +protected def closure (K : convex_cone ℝ E) : convex_cone ℝ E := +{ carrier := closure ↑K, + smul_mem' := + λ c hc _ h₁, map_mem_closure (continuous_id'.const_smul c) h₁ (λ _ h₂, K.smul_mem hc h₂), + add_mem' := λ _ h₁ _ h₂, map_mem_closure₂ continuous_add h₁ h₂ K.add_mem } + +@[simp, norm_cast] lemma coe_closure (K : convex_cone ℝ E) : (K.closure : set E) = closure K := rfl + +protected lemma mem_closure {K : convex_cone ℝ E} {a : E} : + a ∈ K.closure ↔ a ∈ closure (K : set E) := +iff.rfl + +end convex_cone From fecd3520d2a236856f254f27714b80dcfe28ea57 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Tue, 2 May 2023 01:48:44 +0000 Subject: [PATCH 0922/1291] feat(ring_theory/norm): add norm_norm (#18860) Prove the transitivity of the algebra.norm ```lean lemma algebra.norm_norm [is_scalar_tower K L F] [finite_dimensional K F] [is_separable K F] (x : F) : norm K (norm L x) = norm K x ``` Co-authored-by: Eric Wieser --- src/ring_theory/norm.lean | 46 ++++++++++++++++++++++++++++++++++++++- 1 file changed, 45 insertions(+), 1 deletion(-) diff --git a/src/ring_theory/norm.lean b/src/ring_theory/norm.lean index 8c1c231ea7500..d8637f103a938 100644 --- a/src/ring_theory/norm.lean +++ b/src/ring_theory/norm.lean @@ -70,6 +70,14 @@ by { rw [norm_apply, linear_map.det], split_ifs with h, refl } variables {R} +lemma norm_eq_one_of_not_module_finite (h : ¬ module.finite R S) (x : S) : + norm R x = 1 := +begin + refine norm_eq_one_of_not_exists_basis _ (mt _ h) _, + rintro ⟨s, ⟨b⟩⟩, + exact module.finite.of_basis b, +end + -- Can't be a `simp` lemma because it depends on a choice of basis lemma norm_eq_matrix_det [fintype ι] [decidable_eq ι] (b : basis ι R S) (s : S) : norm R s = matrix.det (algebra.left_mul_matrix b s) := @@ -286,7 +294,7 @@ variable (K) of `K`, the norm (down to `K`) of an element `x` of `L` is equal to the product of the images of `x` over all the `K`-embeddings `σ` of `L` into `E`. -/ lemma norm_eq_prod_embeddings [finite_dimensional K L] [is_separable K L] [is_alg_closed E] - {x : L} : algebra_map K E (norm K x) = ∏ σ : L →ₐ[K] E, σ x := + (x : L) : algebra_map K E (norm K x) = ∏ σ : L →ₐ[K] E, σ x := begin have hx := is_separable.is_integral K x, rw [norm_eq_norm_adjoin K x, ring_hom.map_pow, ← adjoin.power_basis_gen hx, @@ -326,6 +334,42 @@ begin { apply_instance } end +variables {F} (L) + +-- TODO. Generalize this proof to rings +lemma norm_norm [algebra L F] [is_scalar_tower K L F] [is_separable K F] (x : F) : + norm K (norm L x) = norm K x := +begin + by_cases hKF : finite_dimensional K F, + { haveI := hKF, + let A := algebraic_closure K, + apply (algebra_map K A).injective, + haveI : finite_dimensional L F := finite_dimensional.right K L F, + haveI : finite_dimensional K L := finite_dimensional.left K L F, + haveI : is_separable K L := is_separable_tower_bot_of_is_separable K L F, + haveI : is_separable L F := is_separable_tower_top_of_is_separable K L F, + letI : ∀ (σ : L →ₐ[K] A), by haveI := σ.to_ring_hom.to_algebra; exact fintype (F →ₐ[L] A) := + λ _, infer_instance, + rw [norm_eq_prod_embeddings K A (_ : F), fintype.prod_equiv alg_hom_equiv_sigma + (λ σ : F →ₐ[K] A, σ x) (λ π : (Σ f : L →ₐ[K] A, _), (π.2 : F → A) x) (λ _, rfl)], + suffices : ∀ σ : L →ₐ[K] A, + by haveI := σ.to_ring_hom.to_algebra; exact ∏ π : F →ₐ[L] A, π x = σ (norm L x), + { simp_rw [← finset.univ_sigma_univ, finset.prod_sigma, this, norm_eq_prod_embeddings], }, + { intro σ, + letI : algebra L A := σ.to_ring_hom.to_algebra, + rw ← norm_eq_prod_embeddings L A (_ : F), + refl, }}, + { rw norm_eq_one_of_not_module_finite hKF, + by_cases hKL : finite_dimensional K L, + { have hLF : ¬ finite_dimensional L F, + { refine (mt _) hKF, + introI hKF, + exact finite_dimensional.trans K L F }, + rw [norm_eq_one_of_not_module_finite hLF, _root_.map_one], }, + { rw norm_eq_one_of_not_module_finite hKL, }}, +end + + end eq_prod_embeddings end algebra From ef74e2bd3a553d31fdb139188a251509b6c6b038 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 2 May 2023 01:48:45 +0000 Subject: [PATCH 0923/1291] refactor(number_theory/modular_forms/slash_actions): use the same notation everywhere (#18886) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Previously we had both `f ∣[k] A` and `f ∣[k, A]` notation, this keeps only the former (on the assumption it more closely resembles the LaTeX). `f ∣[k;γ] A` is then added as notation for the more general `γ`-invariant case that the typeclass is designed around, even though that is never used. This makes `slash_action.map` the canonical spelling of slash actions, as opposed to a mix of this spelling and `modular_form.slash`. Since `modular_form.slash` is no longer canonical, all the lemmas about it are made private as they are immediately subsumed by the typeclass lemmas. Rather than defining the same local notation in each file, this now just opens the locale. This also ungeneralizes `has_zero` + `has_add` to `add_monoid`, since the extra generality is not used, and we do not have it for the (suspiciously!) analogous `distrib_mul_action`. --- src/number_theory/modular_forms/basic.lean | 10 +-- .../modular_forms/slash_actions.lean | 85 +++++++++++++------ .../modular_forms/slash_invariant_forms.lean | 10 +-- 3 files changed, 66 insertions(+), 39 deletions(-) diff --git a/src/number_theory/modular_forms/basic.lean b/src/number_theory/modular_forms/basic.lean index c83217272d5e5..08a2663bcdfff 100644 --- a/src/number_theory/modular_forms/basic.lean +++ b/src/number_theory/modular_forms/basic.lean @@ -35,14 +35,14 @@ open modular_form variables (F : Type*) (Γ : subgroup SL(2, ℤ)) (k : ℤ) -local notation f `∣[`:73 k:0, A `]` :72 := slash_action.map ℂ k A f +open_locale modular_form set_option old_structure_cmd true /--These are `slash_invariant_form`'s that are holomophic and bounded at infinity. -/ structure modular_form extends slash_invariant_form Γ k := (holo' : mdifferentiable 𝓘(ℂ) 𝓘(ℂ) (to_fun : ℍ → ℂ)) -(bdd_at_infty' : ∀ (A : SL(2, ℤ)), is_bounded_at_im_infty (to_fun ∣[k, A])) +(bdd_at_infty' : ∀ (A : SL(2, ℤ)), is_bounded_at_im_infty (to_fun ∣[k] A)) /-- The `slash_invariant_form` associated to a `modular_form`. -/ add_decl_doc modular_form.to_slash_invariant_form @@ -50,7 +50,7 @@ add_decl_doc modular_form.to_slash_invariant_form /--These are `slash_invariant_form`s that are holomophic and zero at infinity. -/ structure cusp_form extends slash_invariant_form Γ k := (holo' : mdifferentiable 𝓘(ℂ) 𝓘(ℂ) (to_fun : ℍ → ℂ)) -(zero_at_infty' : ∀ (A : SL(2, ℤ)), is_zero_at_im_infty (to_fun ∣[k, A])) +(zero_at_infty' : ∀ (A : SL(2, ℤ)), is_zero_at_im_infty (to_fun ∣[k] A)) /-- The `slash_invariant_form` associated to a `cusp_form`. -/ add_decl_doc cusp_form.to_slash_invariant_form @@ -60,14 +60,14 @@ add_decl_doc cusp_form.to_slash_invariant_form at infinity. -/ class modular_form_class extends slash_invariant_form_class F Γ k := (holo: ∀ f : F, mdifferentiable 𝓘(ℂ) 𝓘(ℂ) (f : ℍ → ℂ)) -(bdd_at_infty : ∀ (f : F) (A : SL(2, ℤ)), is_bounded_at_im_infty (f ∣[k, A])) +(bdd_at_infty : ∀ (f : F) (A : SL(2, ℤ)), is_bounded_at_im_infty (f ∣[k] A)) /--`cusp_form_class F Γ k` says that `F` is a type of bundled functions that extend `slash_invariant_forms_class` by requiring that the functions be holomorphic and zero at infinity. -/ class cusp_form_class extends slash_invariant_form_class F Γ k := (holo: ∀ f : F, mdifferentiable 𝓘(ℂ) 𝓘(ℂ) (f : ℍ → ℂ)) -(zero_at_infty : ∀ (f : F) (A : SL(2, ℤ)), is_zero_at_im_infty (f ∣[k, A])) +(zero_at_infty : ∀ (f : F) (A : SL(2, ℤ)), is_zero_at_im_infty (f ∣[k] A)) @[priority 100] instance modular_form_class.modular_form : modular_form_class (modular_form Γ k) Γ k := diff --git a/src/number_theory/modular_forms/slash_actions.lean b/src/number_theory/modular_forms/slash_actions.lean index 472831264b5f9..e03cf93f4925d 100644 --- a/src/number_theory/modular_forms/slash_actions.lean +++ b/src/number_theory/modular_forms/slash_actions.lean @@ -12,6 +12,13 @@ import linear_algebra.matrix.special_linear_group This file defines a class of slash actions, which are families of right actions of a given group parametrized by some Type. This is modeled on the slash action of `GL_pos (fin 2) ℝ` on the space of modular forms. + +## Notation + +In the `modular_form` locale, this provides + +* `f ∣[k;γ] A`: the `k`th `γ`-compatible slash action by `A` on `f` +* `f ∣[k] A`: the `k`th `ℂ`-compatible slash action by `A` on `f`; a shorthand for `f ∣[k;ℂ] A` -/ open complex upper_half_plane @@ -28,7 +35,7 @@ local notation `GL(` n `, ` R `)`⁺ := matrix.GL_pos (fin n) R local notation `SL(` n `, ` R `)` := matrix.special_linear_group (fin n) R /--A general version of the slash action of the space of modular forms.-/ -class slash_action (β G α γ : Type*) [group G] [has_zero α] [has_smul γ α] [has_add α] := +class slash_action (β G α γ : Type*) [group G] [add_monoid α] [has_smul γ α] := (map : β → G → α → α) (zero_slash : ∀ (k : β) (g : G), map k g 0 = 0) (slash_one : ∀ (k : β) (a : α) , map k 1 a = a) @@ -36,8 +43,30 @@ class slash_action (β G α γ : Type*) [group G] [has_zero α] [has_smul γ α] (smul_action : ∀ (k : β) (g : G) (a : α) (z : γ), map k g (z • a) = z • (map k g a)) (add_action : ∀ (k : β) (g : G) (a b : α), map k g (a + b) = map k g a + map k g b) +localized "notation (name := modular_form.slash) f ` ∣[`:100 k `;` γ `] `:0 a :100 := + slash_action.map γ k a f" in modular_form + +localized "notation (name := modular_form.slash_complex) f ` ∣[`:100 k `] `:0 a :100 := + slash_action.map ℂ k a f" in modular_form + +@[simp] lemma slash_action.neg_slash {β G α γ : Type*} [group G] [add_group α] [has_smul γ α] + [slash_action β G α γ] (k : β) (g : G) (a : α) : + (-a) ∣[k;γ] g = - (a ∣[k;γ] g) := +eq_neg_of_add_eq_zero_left $ by rw [←slash_action.add_action, add_left_neg, slash_action.zero_slash] + +@[simp] lemma slash_action.smul_slash_of_tower {R β G α : Type*} (γ : Type*) [group G] [add_group α] + [monoid γ] [mul_action γ α] + [has_smul R γ] [has_smul R α] [is_scalar_tower R γ α] + [slash_action β G α γ] (k : β) (g : G) (a : α) (r : R) : + (r • a) ∣[k;γ] g = r • (a ∣[k;γ] g) := +by rw [←smul_one_smul γ r a, slash_action.smul_action, smul_one_smul] + +attribute [simp] + slash_action.zero_slash slash_action.slash_one + slash_action.smul_action slash_action.add_action + /--Slash_action induced by a monoid homomorphism.-/ -def monoid_hom_slash_action {β G H α γ : Type*} [group G] [has_zero α] [has_smul γ α] [has_add α] +def monoid_hom_slash_action {β G H α γ : Type*} [group G] [add_monoid α] [has_smul γ α] [group H] [slash_action β G α γ] (h : H →* G) : slash_action β H α γ := { map := λ k g, slash_action.map γ k (h g), zero_slash := λ k g, slash_action.zero_slash k (h g), @@ -56,10 +85,12 @@ f (γ • x) * (((↑ₘ γ).det) : ℝ)^(k-1) * (upper_half_plane.denom γ x)^( variables {Γ : subgroup SL(2, ℤ)} {k: ℤ} (f : ℍ → ℂ) -localized "notation (name := modular_form.slash) f ` ∣[`:100 k `]`:0 γ :100 := - modular_form.slash k γ f" in modular_form +section -lemma slash_right_action (k : ℤ) (A B : GL(2, ℝ)⁺) (f : ℍ → ℂ) : +-- temporary notation until the instance is built +local notation f ` ∣[`:100 k `]`:0 γ :100 := modular_form.slash k γ f + +private lemma slash_right_action (k : ℤ) (A B : GL(2, ℝ)⁺) (f : ℍ → ℂ) : (f ∣[k] A) ∣[k] B = f ∣[k] (A * B) := begin ext1, @@ -78,7 +109,7 @@ begin simp_rw [this, ← mul_assoc, ←mul_zpow], end -@[simp] lemma slash_add (k : ℤ) (A : GL(2, ℝ)⁺) (f g : ℍ → ℂ) : +private lemma slash_add (k : ℤ) (A : GL(2, ℝ)⁺) (f g : ℍ → ℂ) : (f + g) ∣[k] A = (f ∣[k] A) + (g ∣[k] A) := begin ext1, @@ -86,12 +117,12 @@ begin ring, end -@[simp] lemma slash_one (k : ℤ) (f : ℍ → ℂ) : (f ∣[k] 1) = f := +private lemma slash_one (k : ℤ) (f : ℍ → ℂ) : (f ∣[k] 1) = f := funext $ by simp [slash] variables {α : Type*} [has_smul α ℂ] [is_scalar_tower α ℂ ℂ] -@[simp] lemma smul_slash (k : ℤ) (A : GL(2, ℝ)⁺) (f : ℍ → ℂ) (c : α) : +private lemma smul_slash (k : ℤ) (A : GL(2, ℝ)⁺) (f : ℍ → ℂ) (c : α) : (c • f) ∣[k] A = c • (f ∣[k] A) := begin simp_rw [←smul_one_smul ℂ c f, ←smul_one_smul ℂ c (f ∣[k] A)], @@ -102,10 +133,7 @@ begin ring, end -@[simp] lemma neg_slash (k : ℤ) (A : GL(2, ℝ)⁺) (f : ℍ → ℂ) : (-f) ∣[k] A = - (f ∣[k] A) := -funext $ by simp [slash] - -@[simp] lemma zero_slash (k : ℤ) (A : GL(2, ℝ)⁺) : (0 : ℍ → ℂ) ∣[k] A = 0 := +private lemma zero_slash (k : ℤ) (A : GL(2, ℝ)⁺) : (0 : ℍ → ℂ) ∣[k] A = 0 := funext $ λ _, by simp only [slash, pi.zero_apply, zero_mul] instance : slash_action ℤ GL(2, ℝ)⁺ (ℍ → ℂ) ℂ := @@ -116,31 +144,32 @@ instance : slash_action ℤ GL(2, ℝ)⁺ (ℍ → ℂ) ℂ := smul_action := smul_slash, add_action := slash_add } +end + +lemma slash_def (A : GL(2, ℝ)⁺) : f ∣[k] A = slash k A f := rfl + instance subgroup_action (Γ : subgroup SL(2, ℤ)) : slash_action ℤ Γ (ℍ → ℂ) ℂ := monoid_hom_slash_action (monoid_hom.comp (matrix.special_linear_group.to_GL_pos) (monoid_hom.comp (matrix.special_linear_group.map (int.cast_ring_hom ℝ)) (subgroup.subtype Γ))) @[simp] lemma subgroup_slash (Γ : subgroup SL(2, ℤ)) (γ : Γ): - (slash_action.map ℂ k γ f) = slash k (γ : GL(2,ℝ)⁺) f := rfl + (f ∣[k] γ) = f ∣[k] (γ : GL(2,ℝ)⁺) := rfl instance SL_action : slash_action ℤ SL(2, ℤ) (ℍ → ℂ) ℂ := monoid_hom_slash_action (monoid_hom.comp (matrix.special_linear_group.to_GL_pos) (matrix.special_linear_group.map (int.cast_ring_hom ℝ))) -@[simp] lemma SL_slash (γ : SL(2, ℤ)): - (slash_action.map ℂ k γ f) = slash k (γ : GL(2,ℝ)⁺) f := rfl - -local notation f `∣[`:73 k:0, A `]` :72 := slash_action.map ℂ k A f +@[simp] lemma SL_slash (γ : SL(2, ℤ)): f ∣[k] γ = f ∣[k] (γ : GL(2,ℝ)⁺) := rfl /-- The constant function 1 is invariant under any element of `SL(2, ℤ)`. -/ -@[simp] lemma is_invariant_one (A : SL(2, ℤ)) : (1 : ℍ → ℂ) ∣[(0 : ℤ), A] = (1 : ℍ → ℂ) := +@[simp] lemma is_invariant_one (A : SL(2, ℤ)) : (1 : ℍ → ℂ) ∣[(0 : ℤ)] A = (1 : ℍ → ℂ) := begin have : (((↑ₘ(A : GL(2,ℝ)⁺)).det) : ℝ) = 1, { simp only [coe_coe, matrix.special_linear_group.coe_GL_pos_coe_GL_coe_matrix, matrix.special_linear_group.det_coe], }, funext, - rw [SL_slash, slash, zero_sub, this], + rw [SL_slash, slash_def, slash, zero_sub, this], simp, end @@ -148,9 +177,9 @@ end if for every matrix `γ ∈ Γ` we have `f(γ • z)= (c*z+d)^k f(z)` where `γ= ![![a, b], ![c, d]]`, and it acts on `ℍ` via Möbius transformations. -/ lemma slash_action_eq'_iff (k : ℤ) (Γ : subgroup SL(2, ℤ)) (f : ℍ → ℂ) (γ : Γ) (z : ℍ) : - f ∣[k, γ] z = f z ↔ f (γ • z) = ((↑ₘ[ℤ]γ 1 0 : ℂ) * z + (↑ₘ[ℤ]γ 1 1 : ℂ))^k * f z := + (f ∣[k] γ) z = f z ↔ f (γ • z) = ((↑ₘ[ℤ]γ 1 0 : ℂ) * z + (↑ₘ[ℤ]γ 1 1 : ℂ))^k * f z := begin - simp only [subgroup_slash, modular_form.slash], + simp only [subgroup_slash, slash_def, modular_form.slash], convert inv_mul_eq_iff_eq_mul₀ _ using 2, { rw mul_comm, simp only [denom, coe_coe, matrix.special_linear_group.coe_GL_pos_coe_GL_coe_matrix, zpow_neg, @@ -161,10 +190,10 @@ begin end lemma mul_slash (k1 k2 : ℤ) (A : GL(2, ℝ)⁺) (f g : ℍ → ℂ) : - (f * g) ∣[k1 + k2, A] = (((↑ₘ A).det) : ℝ) • (f ∣[k1, A]) * (g ∣[k2, A]) := + (f * g) ∣[k1 + k2] A = (((↑ₘ A).det) : ℝ) • (f ∣[k1] A) * (g ∣[k2] A) := begin ext1, - simp only [slash_action.map, slash, matrix.general_linear_group.coe_det_apply, subtype.val_eq_coe, + simp only [slash_def, slash, matrix.general_linear_group.coe_det_apply, subtype.val_eq_coe, pi.mul_apply, pi.smul_apply, algebra.smul_mul_assoc, real_smul], set d : ℂ := ↑((↑ₘ A).det : ℝ), have h1 : d ^ (k1 + k2 - 1) = d * d ^ (k1 - 1) * d ^ (k2 - 1), @@ -182,13 +211,13 @@ begin end @[simp] lemma mul_slash_SL2 (k1 k2 : ℤ) (A : SL(2, ℤ)) (f g : ℍ → ℂ) : - (f * g) ∣[k1 + k2, A] = (f ∣[k1, A]) * (g ∣[k2, A]) := -calc (f * g) ∣[k1 + k2, (A : GL(2, ℝ)⁺)] = _ • (f ∣[k1, A]) * (g ∣[k2, A]) : mul_slash _ _ _ _ _ -... = (1:ℝ) • (f ∣[k1, A]) * (g ∣[k2, A]) : by simp [-matrix.special_linear_group.coe_matrix_coe] -... = (f ∣[k1, A]) * (g ∣[k2, A]) : by simp + (f * g) ∣[k1 + k2] A = (f ∣[k1] A) * (g ∣[k2] A) := +calc (f * g) ∣[k1 + k2] (A : GL(2, ℝ)⁺) = _ • (f ∣[k1] A) * (g ∣[k2] A) : mul_slash _ _ _ _ _ +... = (1:ℝ) • (f ∣[k1] A) * (g ∣[k2] A) : by simp [-matrix.special_linear_group.coe_matrix_coe] +... = (f ∣[k1] A) * (g ∣[k2] A) : by simp lemma mul_slash_subgroup (k1 k2 : ℤ) (Γ : subgroup SL(2, ℤ)) (A : Γ) (f g : ℍ → ℂ) : - (f * g) ∣[k1 + k2, A] = (f ∣[k1, A]) * (g ∣[k2, A]) := + (f * g) ∣[k1 + k2] A = (f ∣[k1] A) * (g ∣[k2] A) := mul_slash_SL2 k1 k2 A f g end modular_form diff --git a/src/number_theory/modular_forms/slash_invariant_forms.lean b/src/number_theory/modular_forms/slash_invariant_forms.lean index 4954d52ee88b9..7ed80a3f81de8 100644 --- a/src/number_theory/modular_forms/slash_invariant_forms.lean +++ b/src/number_theory/modular_forms/slash_invariant_forms.lean @@ -15,7 +15,7 @@ that they form a module. open complex upper_half_plane -open_locale upper_half_plane +open_locale upper_half_plane modular_form noncomputable theory @@ -36,17 +36,15 @@ open modular_form variables (F : Type*) (Γ : out_param $ subgroup SL(2, ℤ)) (k : out_param ℤ) -localized "notation f `∣[`:73 k:0, A `]` :72 := slash_action.map ℂ k A f" in slash_invariant_forms - /--Functions `ℍ → ℂ` that are invariant under the `slash_action`. -/ structure slash_invariant_form := (to_fun : ℍ → ℂ) -(slash_action_eq' : ∀ γ : Γ, to_fun ∣[k, γ] = to_fun) +(slash_action_eq' : ∀ γ : Γ, to_fun ∣[k] γ = to_fun) /--`slash_invariant_form_class F Γ k` asserts `F` is a type of bundled functions that are invariant under the `slash_action`. -/ class slash_invariant_form_class extends fun_like F ℍ (λ _, ℂ) := -(slash_action_eq : ∀ (f : F) (γ : Γ), (f : ℍ → ℂ) ∣[k, γ] = f) +(slash_action_eq : ∀ (f : F) (γ : Γ), (f : ℍ → ℂ) ∣[k] γ = f) attribute [nolint dangerous_instance] slash_invariant_form_class.to_fun_like @@ -135,7 +133,7 @@ instance has_neg : has_neg (slash_invariant_form Γ k) := ⟨ λ f, { to_fun := -f, slash_action_eq' := λ γ, by simpa [modular_form.subgroup_slash, - modular_form.neg_slash] using f.slash_action_eq' γ } ⟩ + slash_action.neg_slash] using f.slash_action_eq' γ } ⟩ @[simp] lemma coe_neg (f : slash_invariant_form Γ k) : ⇑(-f) = -f := rfl @[simp] lemma neg_apply (f : slash_invariant_form Γ k) (z : ℍ) : (-f) z = - (f z) := rfl From 362c2263e25ed3b9ed693773f32f91243612e1da Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Tue, 2 May 2023 01:48:46 +0000 Subject: [PATCH 0924/1291] refactor(category_theory): removing additive_category (#18903) This PR removes the class `additive_category` which was essentially unused, and was creating instance issues during the port. --- src/category_theory/additive/basic.lean | 38 ------------------- .../idempotents/biproducts.lean | 9 ++--- 2 files changed, 3 insertions(+), 44 deletions(-) delete mode 100644 src/category_theory/additive/basic.lean diff --git a/src/category_theory/additive/basic.lean b/src/category_theory/additive/basic.lean deleted file mode 100644 index 0a8dbe2dba74d..0000000000000 --- a/src/category_theory/additive/basic.lean +++ /dev/null @@ -1,38 +0,0 @@ -/- -Copyright (c) 2021 Luke Kershaw. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Luke Kershaw --/ -import category_theory.preadditive.basic -import category_theory.limits.shapes.biproducts - -/-! -# Additive Categories - -This file contains the definition of additive categories. - -TODO: show that finite biproducts implies enriched over commutative monoids and what is missing -additionally to have additivity is that identities have additive inverses, -see https://ncatlab.org/nlab/show/biproduct#BiproductsImplyEnrichment --/ - -noncomputable theory - -open category_theory -open category_theory.preadditive -open category_theory.limits - -universes v v₀ v₁ v₂ u u₀ u₁ u₂ - -namespace category_theory - -variables (C : Type u) [category C] - - -/-- -A preadditive category `C` is called additive if it has all finite biproducts. -See . --/ -class additive_category extends preadditive C, has_finite_biproducts C - -end category_theory diff --git a/src/category_theory/idempotents/biproducts.lean b/src/category_theory/idempotents/biproducts.lean index c91026bedf70b..0041beec748fe 100644 --- a/src/category_theory/idempotents/biproducts.lean +++ b/src/category_theory/idempotents/biproducts.lean @@ -5,14 +5,13 @@ Authors: Joël Riou -/ import category_theory.idempotents.karoubi -import category_theory.additive.basic /-! # Biproducts in the idempotent completion of a preadditive category -In this file, we define an instance expressing that if `C` is an additive category, -then `karoubi C` is also an additive category. +In this file, we define an instance expressing that if `C` is an additive category +(i.e. is preadditive and has finite biproducts), then `karoubi C` is also an additive category. We also obtain that for all `P : karoubi C` where `C` is a preadditive category `C`, there is a canonical isomorphism `P ⊞ P.complement ≅ (to_karoubi C).obj P.X` in the category @@ -97,9 +96,7 @@ lemma karoubi_has_finite_biproducts [has_finite_biproducts C] : simp only [eq_to_hom_refl, id_comp, (F j).idem], }, end, } } -instance {D : Type*} [category D] [additive_category D] : additive_category (karoubi D) := -{ to_preadditive := infer_instance, - to_has_finite_biproducts := karoubi_has_finite_biproducts } +attribute [instance] karoubi_has_finite_biproducts /-- `P.complement` is the formal direct factor of `P.X` given by the idempotent endomorphism `𝟙 P.X - P.p` -/ From 33085c9739c41428651ac461a323fde9a2688d9b Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Tue, 2 May 2023 01:48:47 +0000 Subject: [PATCH 0925/1291] chore(category_theory/monoidal/Mod): rename to Mod_ to avoid name clash in mathlib4 (#18911) This matches the existing `Mon_`. Co-authored-by: Scott Morrison --- .../monoidal/{Mod.lean => Mod_.lean} | 38 +++++++++---------- 1 file changed, 19 insertions(+), 19 deletions(-) rename src/category_theory/monoidal/{Mod.lean => Mod_.lean} (78%) diff --git a/src/category_theory/monoidal/Mod.lean b/src/category_theory/monoidal/Mod_.lean similarity index 78% rename from src/category_theory/monoidal/Mod.lean rename to src/category_theory/monoidal/Mod_.lean index 1ac72705788c9..e88e0ba337171 100644 --- a/src/category_theory/monoidal/Mod.lean +++ b/src/category_theory/monoidal/Mod_.lean @@ -19,26 +19,26 @@ variables (C : Type u₁) [category.{v₁} C] [monoidal_category.{v₁} C] variables {C} /-- A module object for a monoid object, all internal to some monoidal category. -/ -structure Mod (A : Mon_ C) := +structure Mod_ (A : Mon_ C) := (X : C) (act : A.X ⊗ X ⟶ X) (one_act' : (A.one ⊗ 𝟙 X) ≫ act = (λ_ X).hom . obviously) (assoc' : (A.mul ⊗ 𝟙 X) ≫ act = (α_ A.X A.X X).hom ≫ (𝟙 A.X ⊗ act) ≫ act . obviously) -restate_axiom Mod.one_act' -restate_axiom Mod.assoc' -attribute [simp, reassoc] Mod.one_act Mod.assoc +restate_axiom Mod_.one_act' +restate_axiom Mod_.assoc' +attribute [simp, reassoc] Mod_.one_act Mod_.assoc -namespace Mod +namespace Mod_ -variables {A : Mon_ C} (M : Mod A) +variables {A : Mon_ C} (M : Mod_ A) lemma assoc_flip : (𝟙 A.X ⊗ M.act) ≫ M.act = (α_ A.X A.X M.X).inv ≫ (A.mul ⊗ 𝟙 M.X) ≫ M.act := by simp /-- A morphism of module objects. -/ @[ext] -structure hom (M N : Mod A) := +structure hom (M N : Mod_ A) := (hom : M.X ⟶ N.X) (act_hom' : M.act ≫ hom = (𝟙 A.X ⊗ hom) ≫ N.act . obviously) @@ -47,37 +47,37 @@ attribute [simp, reassoc] hom.act_hom /-- The identity morphism on a module object. -/ @[simps] -def id (M : Mod A) : hom M M := +def id (M : Mod_ A) : hom M M := { hom := 𝟙 M.X, } -instance hom_inhabited (M : Mod A) : inhabited (hom M M) := ⟨id M⟩ +instance hom_inhabited (M : Mod_ A) : inhabited (hom M M) := ⟨id M⟩ /-- Composition of module object morphisms. -/ @[simps] -def comp {M N O : Mod A} (f : hom M N) (g : hom N O) : hom M O := +def comp {M N O : Mod_ A} (f : hom M N) (g : hom N O) : hom M O := { hom := f.hom ≫ g.hom, } -instance : category (Mod A) := +instance : category (Mod_ A) := { hom := λ M N, hom M N, id := id, comp := λ M N O f g, comp f g, } -@[simp] lemma id_hom' (M : Mod A) : (𝟙 M : hom M M).hom = 𝟙 M.X := rfl -@[simp] lemma comp_hom' {M N K : Mod A} (f : M ⟶ N) (g : N ⟶ K) : +@[simp] lemma id_hom' (M : Mod_ A) : (𝟙 M : hom M M).hom = 𝟙 M.X := rfl +@[simp] lemma comp_hom' {M N K : Mod_ A} (f : M ⟶ N) (g : N ⟶ K) : (f ≫ g : hom M K).hom = f.hom ≫ g.hom := rfl variables (A) /-- A monoid object as a module over itself. -/ @[simps] -def regular : Mod A := +def regular : Mod_ A := { X := A.X, act := A.mul, } -instance : inhabited (Mod A) := ⟨regular A⟩ +instance : inhabited (Mod_ A) := ⟨regular A⟩ /-- The forgetful functor from module objects to the ambient category. -/ -def forget : Mod A ⥤ C := +def forget : Mod_ A ⥤ C := { obj := λ A, A.X, map := λ A B f, f.hom, } @@ -88,7 +88,7 @@ A morphism of monoid objects induces a "restriction" or "comap" functor between the categories of module objects. -/ @[simps] -def comap {A B : Mon_ C} (f : A ⟶ B) : Mod B ⥤ Mod A := +def comap {A B : Mon_ C} (f : A ⟶ B) : Mod_ B ⥤ Mod_ A := { obj := λ M, { X := M.X, act := (f.hom ⊗ 𝟙 M.X) ≫ M.act, @@ -102,7 +102,7 @@ def comap {A B : Mon_ C} (f : A ⟶ B) : Mod B ⥤ Mod A := -- oh, for homotopy.io in a widget! slice_rhs 2 3 { rw [id_tensor_comp_tensor_id, ←tensor_id_comp_id_tensor], }, rw id_tensor_comp, - slice_rhs 4 5 { rw Mod.assoc_flip, }, + slice_rhs 4 5 { rw Mod_.assoc_flip, }, slice_rhs 3 4 { rw associator_inv_naturality, }, slice_rhs 2 3 { rw [←tensor_id, associator_inv_naturality], }, slice_rhs 1 3 { rw [iso.hom_inv_id_assoc], }, @@ -123,4 +123,4 @@ def comap {A B : Mon_ C} (f : A ⟶ B) : Mod B ⥤ Mod A := -- Lots more could be said about `comap`, e.g. how it interacts with -- identities, compositions, and equalities of monoid object morphisms. -end Mod +end Mod_ From 2c1d8ca2812b64f88992a5294ea3dba144755cd1 Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Tue, 2 May 2023 02:56:29 +0000 Subject: [PATCH 0926/1291] chore(analysis/special_functions): Reduce imports in special_functions (#18909) Remove a bunch of imports in `analysis/special_functions` and `analysis/calculus` which are unnecessary by transitivity, and a few more which are genuinely nontrivial (i.e. the import isn't in the transitive closure of the other imports) but nonetheless aren't needed for the file to compile. --- src/analysis/calculus/bump_function_findim.lean | 3 +-- src/analysis/calculus/bump_function_inner.lean | 3 +-- src/analysis/calculus/cont_diff_def.lean | 2 -- src/analysis/calculus/dslope.lean | 1 - src/analysis/calculus/fderiv.lean | 1 - src/analysis/calculus/fderiv_analytic.lean | 4 ++-- src/analysis/calculus/fderiv_measurable.lean | 2 -- src/analysis/calculus/fderiv_symmetric.lean | 2 -- src/analysis/calculus/inverse.lean | 2 -- src/analysis/calculus/iterated_deriv.lean | 2 +- src/analysis/calculus/local_extr.lean | 2 -- src/analysis/calculus/monotone.lean | 2 -- src/analysis/calculus/parametric_integral.lean | 1 - src/analysis/calculus/taylor.lean | 2 -- src/analysis/special_functions/bernstein.lean | 1 - src/analysis/special_functions/complex/arg.lean | 1 - src/analysis/special_functions/exp.lean | 1 - src/analysis/special_functions/exp_deriv.lean | 1 - src/analysis/special_functions/exponential.lean | 1 - src/analysis/special_functions/gamma.lean | 2 -- src/analysis/special_functions/gaussian.lean | 2 -- src/analysis/special_functions/japanese_bracket.lean | 4 ---- src/analysis/special_functions/log/base.lean | 1 - src/analysis/special_functions/log/monotone.lean | 1 - src/analysis/special_functions/non_integrable.lean | 1 - src/analysis/special_functions/stirling.lean | 2 -- src/analysis/special_functions/trigonometric/angle.lean | 1 - src/analysis/special_functions/trigonometric/basic.lean | 1 - src/analysis/special_functions/trigonometric/bounds.lean | 2 +- src/analysis/special_functions/trigonometric/chebyshev.lean | 2 +- src/analysis/special_functions/trigonometric/complex.lean | 1 - .../special_functions/trigonometric/complex_deriv.lean | 1 - src/analysis/special_functions/trigonometric/deriv.lean | 1 - .../special_functions/trigonometric/euler_sine_prod.lean | 3 +-- src/probability/martingale/upcrossing.lean | 1 + 35 files changed, 9 insertions(+), 51 deletions(-) diff --git a/src/analysis/calculus/bump_function_findim.lean b/src/analysis/calculus/bump_function_findim.lean index 66e5922f13c04..e44927acf9dd9 100644 --- a/src/analysis/calculus/bump_function_findim.lean +++ b/src/analysis/calculus/bump_function_findim.lean @@ -3,11 +3,10 @@ Copyright (c) 2022 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ -import analysis.calculus.bump_function_inner import analysis.calculus.series import analysis.convolution +import analysis.inner_product_space.euclidean_dist import data.set.pointwise.support -import measure_theory.measure.haar_lebesgue /-! # Bump functions in finite-dimensional vector spaces diff --git a/src/analysis/calculus/bump_function_inner.lean b/src/analysis/calculus/bump_function_inner.lean index af3a1238fc7a4..f236cbd0c3333 100644 --- a/src/analysis/calculus/bump_function_inner.lean +++ b/src/analysis/calculus/bump_function_inner.lean @@ -4,8 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel, Floris van Doorn -/ import analysis.calculus.iterated_deriv -import analysis.inner_product_space.euclidean_dist -import measure_theory.function.locally_integrable +import analysis.inner_product_space.calculus import measure_theory.integral.set_integral /-! diff --git a/src/analysis/calculus/cont_diff_def.lean b/src/analysis/calculus/cont_diff_def.lean index 8720cd1832c86..b65c429a34abf 100644 --- a/src/analysis/calculus/cont_diff_def.lean +++ b/src/analysis/calculus/cont_diff_def.lean @@ -4,9 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ import analysis.calculus.fderiv -import analysis.normed_space.multilinear import analysis.calculus.formal_multilinear_series -import data.enat.basic /-! # Higher differentiability diff --git a/src/analysis/calculus/dslope.lean b/src/analysis/calculus/dslope.lean index 8899e73779af2..48b5af6941003 100644 --- a/src/analysis/calculus/dslope.lean +++ b/src/analysis/calculus/dslope.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import analysis.calculus.deriv -import linear_algebra.affine_space.slope /-! # Slope of a differentiable function diff --git a/src/analysis/calculus/fderiv.lean b/src/analysis/calculus/fderiv.lean index 5953792f5af23..de89c70222e52 100644 --- a/src/analysis/calculus/fderiv.lean +++ b/src/analysis/calculus/fderiv.lean @@ -6,7 +6,6 @@ Authors: Jeremy Avigad, Sébastien Gouëzel, Yury Kudryashov import analysis.asymptotics.asymptotic_equivalent import analysis.calculus.tangent_cone import analysis.normed_space.bounded_linear_maps -import analysis.normed_space.units /-! # The Fréchet derivative diff --git a/src/analysis/calculus/fderiv_analytic.lean b/src/analysis/calculus/fderiv_analytic.lean index 58483589cadf0..bddec64dbb2bc 100644 --- a/src/analysis/calculus/fderiv_analytic.lean +++ b/src/analysis/calculus/fderiv_analytic.lean @@ -3,9 +3,9 @@ Copyright (c) 2021 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ -import analysis.calculus.deriv import analysis.analytic.basic -import analysis.calculus.cont_diff +import analysis.calculus.deriv +import analysis.calculus.cont_diff_def /-! # Frechet derivatives of analytic functions. diff --git a/src/analysis/calculus/fderiv_measurable.lean b/src/analysis/calculus/fderiv_measurable.lean index 6c52ff3f17c76..51d4260b59aff 100644 --- a/src/analysis/calculus/fderiv_measurable.lean +++ b/src/analysis/calculus/fderiv_measurable.lean @@ -4,9 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel, Yury Kudryashov -/ import analysis.calculus.deriv -import measure_theory.constructions.borel_space import measure_theory.function.strongly_measurable.basic -import tactic.ring_exp /-! # Derivative is measurable diff --git a/src/analysis/calculus/fderiv_symmetric.lean b/src/analysis/calculus/fderiv_symmetric.lean index 70a4aeab6a8ff..52a36b3928461 100644 --- a/src/analysis/calculus/fderiv_symmetric.lean +++ b/src/analysis/calculus/fderiv_symmetric.lean @@ -3,9 +3,7 @@ Copyright (c) 2021 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ -import analysis.calculus.deriv import analysis.calculus.mean_value -import analysis.convex.topology /-! # Symmetry of the second derivative diff --git a/src/analysis/calculus/inverse.lean b/src/analysis/calculus/inverse.lean index 6f928e563362d..de850dabe3232 100644 --- a/src/analysis/calculus/inverse.lean +++ b/src/analysis/calculus/inverse.lean @@ -4,9 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov, Heather Macbeth, Sébastien Gouëzel -/ import analysis.calculus.cont_diff -import tactic.ring_exp import analysis.normed_space.banach -import topology.local_homeomorph /-! # Inverse function theorem diff --git a/src/analysis/calculus/iterated_deriv.lean b/src/analysis/calculus/iterated_deriv.lean index 2d1050774167c..67837508500c8 100644 --- a/src/analysis/calculus/iterated_deriv.lean +++ b/src/analysis/calculus/iterated_deriv.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ import analysis.calculus.deriv -import analysis.calculus.cont_diff +import analysis.calculus.cont_diff_def /-! # One-dimensional iterated derivatives diff --git a/src/analysis/calculus/local_extr.lean b/src/analysis/calculus/local_extr.lean index 70ee18d3fc286..ef494a5da1bcb 100644 --- a/src/analysis/calculus/local_extr.lean +++ b/src/analysis/calculus/local_extr.lean @@ -4,10 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import analysis.calculus.deriv -import data.polynomial.field_division import topology.algebra.order.extend_from import topology.algebra.polynomial -import topology.local_extr /-! # Local extrema of smooth functions diff --git a/src/analysis/calculus/monotone.lean b/src/analysis/calculus/monotone.lean index cabfe29cebc40..baf24c83fbf35 100644 --- a/src/analysis/calculus/monotone.lean +++ b/src/analysis/calculus/monotone.lean @@ -3,8 +3,6 @@ Copyright (c) 2022 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ -import measure_theory.measure.lebesgue -import analysis.calculus.deriv import measure_theory.covering.one_dim import order.monotone.extension diff --git a/src/analysis/calculus/parametric_integral.lean b/src/analysis/calculus/parametric_integral.lean index ea58e3346dfc1..bdbb1d29fc49f 100644 --- a/src/analysis/calculus/parametric_integral.lean +++ b/src/analysis/calculus/parametric_integral.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Patrick Massot -/ import measure_theory.integral.set_integral -import analysis.calculus.mean_value /-! # Derivatives of integrals depending on parameters diff --git a/src/analysis/calculus/taylor.lean b/src/analysis/calculus/taylor.lean index 2ac9483bbf23c..0ad33d0d28309 100644 --- a/src/analysis/calculus/taylor.lean +++ b/src/analysis/calculus/taylor.lean @@ -3,10 +3,8 @@ Copyright (c) 2022 Moritz Doll. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Moritz Doll -/ - import analysis.calculus.iterated_deriv import analysis.calculus.mean_value -import data.polynomial.basic import data.polynomial.module /-! diff --git a/src/analysis/special_functions/bernstein.lean b/src/analysis/special_functions/bernstein.lean index 43cc7cc4cfd3f..e26be5fe94937 100644 --- a/src/analysis/special_functions/bernstein.lean +++ b/src/analysis/special_functions/bernstein.lean @@ -3,7 +3,6 @@ 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.order.field.basic import analysis.specific_limits.basic import ring_theory.polynomial.bernstein import topology.continuous_function.polynomial diff --git a/src/analysis/special_functions/complex/arg.lean b/src/analysis/special_functions/complex/arg.lean index 76fa41d560635..28ee90e5089ed 100644 --- a/src/analysis/special_functions/complex/arg.lean +++ b/src/analysis/special_functions/complex/arg.lean @@ -3,7 +3,6 @@ Copyright (c) 2018 Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes, Abhimanyu Pallavi Sudhir, Jean Lo, Calle Sönne, Benjamin Davidson -/ -import algebra.order.to_interval_mod import analysis.special_functions.trigonometric.angle import analysis.special_functions.trigonometric.inverse diff --git a/src/analysis/special_functions/exp.lean b/src/analysis/special_functions/exp.lean index 0a00ddc6869ed..b1b377aef9ff6 100644 --- a/src/analysis/special_functions/exp.lean +++ b/src/analysis/special_functions/exp.lean @@ -6,7 +6,6 @@ Authors: Chris Hughes, Abhimanyu Pallavi Sudhir, Jean Lo, Calle Sönne import analysis.asymptotics.theta import analysis.complex.basic import analysis.specific_limits.normed -import data.complex.exponential /-! # Complex and real exponential diff --git a/src/analysis/special_functions/exp_deriv.lean b/src/analysis/special_functions/exp_deriv.lean index 8bd19a0140f00..ff0b7ade2a641 100644 --- a/src/analysis/special_functions/exp_deriv.lean +++ b/src/analysis/special_functions/exp_deriv.lean @@ -5,7 +5,6 @@ Authors: Chris Hughes, Abhimanyu Pallavi Sudhir, Jean Lo, Calle Sönne -/ import analysis.calculus.inverse import analysis.complex.real_deriv -import analysis.special_functions.exp /-! # Complex and real exponential diff --git a/src/analysis/special_functions/exponential.lean b/src/analysis/special_functions/exponential.lean index d7e81c0cf58ce..dccb029c003ca 100644 --- a/src/analysis/special_functions/exponential.lean +++ b/src/analysis/special_functions/exponential.lean @@ -5,7 +5,6 @@ Authors: Anatole Dedecker -/ import analysis.normed_space.exponential import analysis.calculus.fderiv_analytic -import data.complex.exponential import topology.metric_space.cau_seq_filter /-! diff --git a/src/analysis/special_functions/gamma.lean b/src/analysis/special_functions/gamma.lean index 964fc1cc8908e..cecff8716528e 100644 --- a/src/analysis/special_functions/gamma.lean +++ b/src/analysis/special_functions/gamma.lean @@ -4,8 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: David Loeffler -/ import measure_theory.integral.exp_decay -import analysis.calculus.parametric_integral -import analysis.special_functions.integrals import analysis.convolution import analysis.special_functions.trigonometric.euler_sine_prod diff --git a/src/analysis/special_functions/gaussian.lean b/src/analysis/special_functions/gaussian.lean index 726f21901ca4a..5825b1f8a1398 100644 --- a/src/analysis/special_functions/gaussian.lean +++ b/src/analysis/special_functions/gaussian.lean @@ -7,9 +7,7 @@ Authors: Sébastien Gouëzel import analysis.special_functions.gamma import analysis.special_functions.polar_coord import analysis.convex.complex -import analysis.normed.group.basic import analysis.complex.cauchy_integral -import measure_theory.group.integration import analysis.fourier.poisson_summation /-! diff --git a/src/analysis/special_functions/japanese_bracket.lean b/src/analysis/special_functions/japanese_bracket.lean index 1ebd195adc4bc..3ed378923a66a 100644 --- a/src/analysis/special_functions/japanese_bracket.lean +++ b/src/analysis/special_functions/japanese_bracket.lean @@ -3,11 +3,7 @@ Copyright (c) 2022 Moritz Doll. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Moritz Doll -/ - -import analysis.special_functions.integrals -import analysis.special_functions.pow import measure_theory.integral.layercake -import tactic.positivity /-! # Japanese Bracket diff --git a/src/analysis/special_functions/log/base.lean b/src/analysis/special_functions/log/base.lean index a4d25840c2688..55ad5d89c3c36 100644 --- a/src/analysis/special_functions/log/base.lean +++ b/src/analysis/special_functions/log/base.lean @@ -3,7 +3,6 @@ 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.basic import analysis.special_functions.pow import data.int.log diff --git a/src/analysis/special_functions/log/monotone.lean b/src/analysis/special_functions/log/monotone.lean index 97e9a4f0d190f..f0981e4f14658 100644 --- a/src/analysis/special_functions/log/monotone.lean +++ b/src/analysis/special_functions/log/monotone.lean @@ -3,7 +3,6 @@ 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 /-! diff --git a/src/analysis/special_functions/non_integrable.lean b/src/analysis/special_functions/non_integrable.lean index a762d8bb77764..d3ae8b75964d5 100644 --- a/src/analysis/special_functions/non_integrable.lean +++ b/src/analysis/special_functions/non_integrable.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import analysis.special_functions.integrals -import analysis.calculus.fderiv_measurable /-! # Non integrable functions diff --git a/src/analysis/special_functions/stirling.lean b/src/analysis/special_functions/stirling.lean index 8e859206a3df4..9852179666596 100644 --- a/src/analysis/special_functions/stirling.lean +++ b/src/analysis/special_functions/stirling.lean @@ -4,8 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Moritz Firsching, Fabian Kruse, Nikolas Kuhn -/ import analysis.p_series -import analysis.special_functions.log.deriv -import tactic.positivity import data.real.pi.wallis /-! diff --git a/src/analysis/special_functions/trigonometric/angle.lean b/src/analysis/special_functions/trigonometric/angle.lean index f7eed78ad9d96..958ca0e974bbb 100644 --- a/src/analysis/special_functions/trigonometric/angle.lean +++ b/src/analysis/special_functions/trigonometric/angle.lean @@ -6,7 +6,6 @@ Authors: Calle Sönne import analysis.special_functions.trigonometric.basic import analysis.normed.group.add_circle import algebra.char_zero.quotient -import algebra.order.to_interval_mod import topology.instances.sign /-! diff --git a/src/analysis/special_functions/trigonometric/basic.lean b/src/analysis/special_functions/trigonometric/basic.lean index f605746cd1007..2473599ce5a67 100644 --- a/src/analysis/special_functions/trigonometric/basic.lean +++ b/src/analysis/special_functions/trigonometric/basic.lean @@ -4,7 +4,6 @@ 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.exp -import data.set.intervals.infinite /-! # Trigonometric functions diff --git a/src/analysis/special_functions/trigonometric/bounds.lean b/src/analysis/special_functions/trigonometric/bounds.lean index e185a6bbf1efe..7e2e8244e4e08 100644 --- a/src/analysis/special_functions/trigonometric/bounds.lean +++ b/src/analysis/special_functions/trigonometric/bounds.lean @@ -3,8 +3,8 @@ Copyright (c) 2022 David Loeffler. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: David Loeffler -/ -import analysis.special_functions.trigonometric.basic import analysis.special_functions.trigonometric.arctan_deriv + /-! # Polynomial bounds for trigonometric functions diff --git a/src/analysis/special_functions/trigonometric/chebyshev.lean b/src/analysis/special_functions/trigonometric/chebyshev.lean index aa8a8e90c8672..ea96d6f5337d2 100644 --- a/src/analysis/special_functions/trigonometric/chebyshev.lean +++ b/src/analysis/special_functions/trigonometric/chebyshev.lean @@ -3,8 +3,8 @@ Copyright (c) 2020 Johan Commelin. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin -/ -import analysis.complex.basic import data.complex.exponential +import data.complex.module import data.polynomial.algebra_map import ring_theory.polynomial.chebyshev diff --git a/src/analysis/special_functions/trigonometric/complex.lean b/src/analysis/special_functions/trigonometric/complex.lean index 2e10bbd8af22a..5291d37b34aac 100644 --- a/src/analysis/special_functions/trigonometric/complex.lean +++ b/src/analysis/special_functions/trigonometric/complex.lean @@ -4,7 +4,6 @@ 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 algebra.quadratic_discriminant -import analysis.special_functions.trigonometric.basic import analysis.convex.specific_functions /-! diff --git a/src/analysis/special_functions/trigonometric/complex_deriv.lean b/src/analysis/special_functions/trigonometric/complex_deriv.lean index 6e2fcb65e2b18..a1c2f4550f6e8 100644 --- a/src/analysis/special_functions/trigonometric/complex_deriv.lean +++ b/src/analysis/special_functions/trigonometric/complex_deriv.lean @@ -4,7 +4,6 @@ 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.trigonometric.complex -import analysis.special_functions.trigonometric.deriv /-! # Complex trigonometric functions diff --git a/src/analysis/special_functions/trigonometric/deriv.lean b/src/analysis/special_functions/trigonometric/deriv.lean index 5edde7a39cbab..29a599126d297 100644 --- a/src/analysis/special_functions/trigonometric/deriv.lean +++ b/src/analysis/special_functions/trigonometric/deriv.lean @@ -6,7 +6,6 @@ Authors: Chris Hughes, Abhimanyu Pallavi Sudhir, Jean Lo, Calle Sönne, Benjamin import order.monotone.odd import analysis.special_functions.exp_deriv import analysis.special_functions.trigonometric.basic -import data.set.intervals.monotone /-! # Differentiability of trigonometric functions diff --git a/src/analysis/special_functions/trigonometric/euler_sine_prod.lean b/src/analysis/special_functions/trigonometric/euler_sine_prod.lean index 902152adc5abf..7fd83358cee89 100644 --- a/src/analysis/special_functions/trigonometric/euler_sine_prod.lean +++ b/src/analysis/special_functions/trigonometric/euler_sine_prod.lean @@ -3,8 +3,7 @@ Copyright (c) 2023 David Loeffler. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: David Loeffler -/ - -import data.real.pi.wallis +import analysis.special_functions.integrals import measure_theory.integral.peak_function /-! # Euler's infinite product for the sine function diff --git a/src/probability/martingale/upcrossing.lean b/src/probability/martingale/upcrossing.lean index df61f68c204c9..2bc98deb4de6a 100644 --- a/src/probability/martingale/upcrossing.lean +++ b/src/probability/martingale/upcrossing.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 Kexing Ying. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kexing Ying -/ +import data.set.intervals.monotone import probability.process.hitting_time import probability.martingale.basic From 11d5ff217c07a8070cddfc5d94608854306f7f68 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Tue, 2 May 2023 11:42:39 +0000 Subject: [PATCH 0927/1291] feat(measure_theory/constructions/polish): a Polish Borel space is measurably equivalent to a subset of the reals (#18881) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: RemyDegenne Co-authored-by: Rémy Degenne --- src/measure_theory/constructions/polish.lean | 69 ++++++++++++++++++-- 1 file changed, 65 insertions(+), 4 deletions(-) diff --git a/src/measure_theory/constructions/polish.lean b/src/measure_theory/constructions/polish.lean index 43bc4770d445f..685cab7f6fd2e 100644 --- a/src/measure_theory/constructions/polish.lean +++ b/src/measure_theory/constructions/polish.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel, Felix Weilacher -/ +import data.real.cardinality import topology.perfect import measure_theory.constructions.borel_space @@ -728,6 +729,8 @@ begin exact this.polish_space, end +namespace polish_space + /-Note: This is to avoid a loop in TC inference. When ported to Lean 4, this will not be necessary, and `second_countable_of_polish` should probably just be added as an instance soon after the definition of `polish_space`.-/ @@ -737,15 +740,12 @@ h.second_countable local attribute [-instance] polish_space_of_complete_second_countable local attribute [instance] second_countable_of_polish -namespace polish_space - variables {β : Type*} [topological_space β] [polish_space α] [polish_space β] variables [measurable_space α] [measurable_space β] [borel_space α] [borel_space β] -noncomputable theory - /-- If two Polish spaces admit Borel measurable injections to one another, then they are Borel isomorphic.-/ +noncomputable def borel_schroeder_bernstein {f : α → β} {g : β → α} (fmeas : measurable f) (finj : function.injective f) @@ -754,6 +754,7 @@ def borel_schroeder_bernstein (fmeas.measurable_embedding finj).schroeder_bernstein (gmeas.measurable_embedding ginj) /-- Any uncountable Polish space is Borel isomorphic to the Cantor space `ℕ → bool`.-/ +noncomputable def measurable_equiv_nat_bool_of_not_countable (h : ¬ countable α) : α ≃ᵐ (ℕ → bool) := begin apply nonempty.some, @@ -764,12 +765,14 @@ begin end /-- The **Borel Isomorphism Theorem**: Any two uncountable Polish spaces are Borel isomorphic.-/ +noncomputable def measurable_equiv_of_not_countable (hα : ¬ countable α) (hβ : ¬ countable β ) : α ≃ᵐ β := (measurable_equiv_nat_bool_of_not_countable hα).trans (measurable_equiv_nat_bool_of_not_countable hβ).symm /-- The **Borel Isomorphism Theorem**: If two Polish spaces have the same cardinality, they are Borel isomorphic.-/ +noncomputable def equiv.measurable_equiv (e : α ≃ β) : α ≃ᵐ β := begin by_cases h : countable α, @@ -781,3 +784,61 @@ begin end end polish_space + + +namespace measure_theory + +-- todo after the port: move to topology/metric_space/polish +instance [polish_space α] : polish_space (univ : set α) := is_closed_univ.polish_space + +variables (α) [measurable_space α] [polish_space α] [borel_space α] + +lemma exists_nat_measurable_equiv_range_coe_fin_of_finite [finite α] : + ∃ n : ℕ, nonempty (α ≃ᵐ range (coe : fin n → ℝ)) := +begin + obtain ⟨n, ⟨n_equiv⟩⟩ := finite.exists_equiv_fin α, + refine ⟨n, ⟨polish_space.equiv.measurable_equiv (n_equiv.trans _)⟩⟩, + exact equiv.of_injective _ (nat.cast_injective.comp fin.val_injective), +end + +lemma measurable_equiv_range_coe_nat_of_infinite_of_countable [infinite α] [countable α] : + nonempty (α ≃ᵐ range (coe : ℕ → ℝ)) := +begin + haveI : polish_space (range (coe : ℕ → ℝ)), + { exact nat.closed_embedding_coe_real.is_closed_map.closed_range.polish_space, }, + refine ⟨polish_space.equiv.measurable_equiv _⟩, + refine (nonempty_equiv_of_countable.some : α ≃ ℕ).trans _, + exact equiv.of_injective coe nat.cast_injective, +end + +/-- Any Polish Borel space is measurably equivalent to a subset of the reals. -/ +theorem exists_subset_real_measurable_equiv : ∃ s : set ℝ, measurable_set s ∧ nonempty (α ≃ᵐ s) := +begin + by_cases hα : countable α, + { casesI finite_or_infinite α, + { obtain ⟨n, h_nonempty_equiv⟩ := exists_nat_measurable_equiv_range_coe_fin_of_finite α, + refine ⟨_, _, h_nonempty_equiv⟩, + letI : measurable_space (fin n) := borel (fin n), + haveI : borel_space (fin n) := ⟨rfl⟩, + refine measurable_embedding.measurable_set_range _, + { apply_instance, }, + { exact continuous_of_discrete_topology.measurable_embedding + (nat.cast_injective.comp fin.val_injective), }, }, + { refine ⟨_, _, measurable_equiv_range_coe_nat_of_infinite_of_countable α⟩, + refine measurable_embedding.measurable_set_range _, + { apply_instance, }, + { exact continuous_of_discrete_topology.measurable_embedding nat.cast_injective, }, }, }, + { refine ⟨univ, measurable_set.univ, + ⟨(polish_space.measurable_equiv_of_not_countable hα _ : α ≃ᵐ (univ : set ℝ))⟩⟩, + rw countable_coe_iff, + exact cardinal.not_countable_real, } +end + +/-- Any Polish Borel space embeds measurably into the reals. -/ +theorem exists_measurable_embedding_real : ∃ (f : α → ℝ), measurable_embedding f := +begin + obtain ⟨s, hs, ⟨e⟩⟩ := exists_subset_real_measurable_equiv α, + exact ⟨coe ∘ e, (measurable_embedding.subtype_coe hs).comp e.measurable_embedding⟩, +end + +end measure_theory From ef55335933293309ff8c0b1d20ffffeecbe5c39f Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Tue, 2 May 2023 13:48:08 +0000 Subject: [PATCH 0928/1291] chore(*): add mathlib4 synchronization comments (#18897) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.category.Group.Z_Module_equivalence` * `algebraic_topology.extra_degeneracy` * `category_theory.abelian.exact` * `category_theory.abelian.transfer` * `category_theory.monad.coequalizer` * `category_theory.preadditive.projective_resolution` * `data.mv_polynomial.polynomial` * `data.ordmap.ordnode` * `field_theory.tower` * `linear_algebra.projective_space.basic` * `linear_algebra.projective_space.independence` * `linear_algebra.symplectic_group` * `topology.category.Profinite.basic` * `topology.category.Profinite.projective` --- src/algebra/category/Group/Z_Module_equivalence.lean | 3 +++ src/algebraic_topology/extra_degeneracy.lean | 3 +++ src/category_theory/abelian/exact.lean | 3 +++ src/category_theory/abelian/transfer.lean | 3 +++ src/category_theory/monad/coequalizer.lean | 3 +++ src/category_theory/preadditive/projective_resolution.lean | 3 +++ src/data/mv_polynomial/polynomial.lean | 3 +++ src/data/ordmap/ordnode.lean | 3 +++ src/field_theory/tower.lean | 3 +++ src/linear_algebra/projective_space/basic.lean | 3 +++ src/linear_algebra/projective_space/independence.lean | 3 +++ src/linear_algebra/symplectic_group.lean | 3 +++ src/topology/category/Profinite/basic.lean | 3 +++ src/topology/category/Profinite/projective.lean | 3 +++ 14 files changed, 42 insertions(+) diff --git a/src/algebra/category/Group/Z_Module_equivalence.lean b/src/algebra/category/Group/Z_Module_equivalence.lean index 5b0097f651da5..d24307095d6e2 100644 --- a/src/algebra/category/Group/Z_Module_equivalence.lean +++ b/src/algebra/category/Group/Z_Module_equivalence.lean @@ -6,6 +6,9 @@ Authors: Scott Morrison import algebra.category.Module.basic /-! +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The forgetful functor from ℤ-modules to additive commutative groups is an equivalence of categories. diff --git a/src/algebraic_topology/extra_degeneracy.lean b/src/algebraic_topology/extra_degeneracy.lean index 0499cbf96e582..fe97e110dc131 100644 --- a/src/algebraic_topology/extra_degeneracy.lean +++ b/src/algebraic_topology/extra_degeneracy.lean @@ -14,6 +14,9 @@ import tactic.fin_cases # Augmented simplicial objects with an extra degeneracy +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In simplicial homotopy theory, in order to prove that the connected components of a simplicial set `X` are contractible, it suffices to construct an extra degeneracy as it is defined in *Simplicial Homotopy Theory* by Goerss-Jardine p. 190. diff --git a/src/category_theory/abelian/exact.lean b/src/category_theory/abelian/exact.lean index 67d1b304d6fdd..596cf13dbbb76 100644 --- a/src/category_theory/abelian/exact.lean +++ b/src/category_theory/abelian/exact.lean @@ -14,6 +14,9 @@ import tactic.tfae /-! # Exact sequences in abelian categories +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In an abelian category, we get several interesting results related to exactness which are not true in more general settings. diff --git a/src/category_theory/abelian/transfer.lean b/src/category_theory/abelian/transfer.lean index 1ce54b6ef630b..5134e852c5df9 100644 --- a/src/category_theory/abelian/transfer.lean +++ b/src/category_theory/abelian/transfer.lean @@ -10,6 +10,9 @@ import category_theory.adjunction.limits /-! # Transferring "abelian-ness" across a functor +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + If `C` is an additive category, `D` is an abelian category, we have `F : C ⥤ D` `G : D ⥤ C` (both preserving zero morphisms), `G` is left exact (that is, preserves finite limits), diff --git a/src/category_theory/monad/coequalizer.lean b/src/category_theory/monad/coequalizer.lean index 8e08c1c3ff8c6..3954e674ec514 100644 --- a/src/category_theory/monad/coequalizer.lean +++ b/src/category_theory/monad/coequalizer.lean @@ -11,6 +11,9 @@ import category_theory.monad.algebra /-! # Special coequalizers associated to a monad +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Associated to a monad `T : C ⥤ C` we have important coequalizer constructions: Any algebra is a coequalizer (in the category of algebras) of free algebras. Furthermore, this coequalizer is reflexive. diff --git a/src/category_theory/preadditive/projective_resolution.lean b/src/category_theory/preadditive/projective_resolution.lean index d04bcf9f2bbb5..16f66b844bf9e 100644 --- a/src/category_theory/preadditive/projective_resolution.lean +++ b/src/category_theory/preadditive/projective_resolution.lean @@ -10,6 +10,9 @@ import algebra.homology.homotopy_category /-! # Projective resolutions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A projective resolution `P : ProjectiveResolution Z` of an object `Z : C` consists of a `ℕ`-indexed chain complex `P.complex` of projective objects, along with a chain map `P.π` from `C` to the chain complex consisting just of `Z` in degree zero, diff --git a/src/data/mv_polynomial/polynomial.lean b/src/data/mv_polynomial/polynomial.lean index d3571c14ddbb5..bbe843c6ba862 100644 --- a/src/data/mv_polynomial/polynomial.lean +++ b/src/data/mv_polynomial/polynomial.lean @@ -8,6 +8,9 @@ import data.polynomial.eval /-! # Some lemmas relating polynomials and multivariable polynomials. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ namespace mv_polynomial diff --git a/src/data/ordmap/ordnode.lean b/src/data/ordmap/ordnode.lean index 3f03aeca20423..66d7d2e99ca78 100644 --- a/src/data/ordmap/ordnode.lean +++ b/src/data/ordmap/ordnode.lean @@ -10,6 +10,9 @@ import data.nat.psub /-! # Ordered sets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines a data structure for ordered sets, supporting a variety of useful operations including insertion and deletion, logarithmic time lookup, set operations, folds, diff --git a/src/field_theory/tower.lean b/src/field_theory/tower.lean index f35f35b430c1d..d3c38ccef9dff 100644 --- a/src/field_theory/tower.lean +++ b/src/field_theory/tower.lean @@ -12,6 +12,9 @@ import linear_algebra.free_module.finite.matrix /-! # Tower of field extensions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove the tower law for arbitrary extensions and finite extensions. Suppose `L` is a field extension of `K` and `K` is a field extension of `F`. Then `[L:F] = [L:K] [K:F]` where `[E₁:E₂]` means the `E₂`-dimension of `E₁`. diff --git a/src/linear_algebra/projective_space/basic.lean b/src/linear_algebra/projective_space/basic.lean index 20cab642407c3..78593079ab34d 100644 --- a/src/linear_algebra/projective_space/basic.lean +++ b/src/linear_algebra/projective_space/basic.lean @@ -9,6 +9,9 @@ import linear_algebra.finite_dimensional # Projective Spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains the definition of the projectivization of a vector space over a field, as well as the bijection between said projectivization and the collection of all one dimensional subspaces of the vector space. diff --git a/src/linear_algebra/projective_space/independence.lean b/src/linear_algebra/projective_space/independence.lean index f210d15c03dc1..201eb9155d050 100644 --- a/src/linear_algebra/projective_space/independence.lean +++ b/src/linear_algebra/projective_space/independence.lean @@ -9,6 +9,9 @@ import linear_algebra.projective_space.basic /-! # Independence in Projective Space +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define independence and dependence of families of elements in projective space. ## Implementation Details diff --git a/src/linear_algebra/symplectic_group.lean b/src/linear_algebra/symplectic_group.lean index 3e04c91cdebde..895a077101aae 100644 --- a/src/linear_algebra/symplectic_group.lean +++ b/src/linear_algebra/symplectic_group.lean @@ -9,6 +9,9 @@ import linear_algebra.matrix.nonsingular_inverse /-! # The Symplectic Group +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the symplectic group and proves elementary properties. ## Main Definitions diff --git a/src/topology/category/Profinite/basic.lean b/src/topology/category/Profinite/basic.lean index b07db2c8952d4..8f71fdef069dc 100644 --- a/src/topology/category/Profinite/basic.lean +++ b/src/topology/category/Profinite/basic.lean @@ -15,6 +15,9 @@ import category_theory.Fintype /-! # The category of Profinite Types +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We construct the category of profinite topological spaces, often called profinite sets -- perhaps they could be called profinite types in Lean. diff --git a/src/topology/category/Profinite/projective.lean b/src/topology/category/Profinite/projective.lean index 8bbb927783a10..64b5bef31b325 100644 --- a/src/topology/category/Profinite/projective.lean +++ b/src/topology/category/Profinite/projective.lean @@ -11,6 +11,9 @@ import category_theory.preadditive.projective /-! # Profinite sets have enough projectives +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we show that `Profinite` has enough projectives. ## Main results From 36b8aa61ea7c05727161f96a0532897bd72aedab Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Tue, 2 May 2023 16:15:17 +0000 Subject: [PATCH 0929/1291] feat(algebra/algebra/basic): add shortcut instance (#18907) --- src/algebra/algebra/basic.lean | 5 +++++ 1 file changed, 5 insertions(+) diff --git a/src/algebra/algebra/basic.lean b/src/algebra/algebra/basic.lean index 261a01004e27b..518dfad84113e 100644 --- a/src/algebra/algebra/basic.lean +++ b/src/algebra/algebra/basic.lean @@ -780,6 +780,11 @@ instance is_scalar_tower.to_smul_comm_class : smul_comm_class R A M := instance is_scalar_tower.to_smul_comm_class' : smul_comm_class A R M := smul_comm_class.symm _ _ _ +@[priority 200] -- see Note [lower instance priority] +instance algebra.to_smul_comm_class {R A} [comm_semiring R] [semiring A] [algebra R A] : + smul_comm_class R A A := +is_scalar_tower.to_smul_comm_class + lemma smul_algebra_smul_comm (r : R) (a : A) (m : M) : a • r • m = r • a • m := smul_comm _ _ _ From 7413128c3bcb3b0818e3e18720abc9ea3100fb49 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Tue, 2 May 2023 21:10:16 +0000 Subject: [PATCH 0930/1291] feat(data/list/cycle): well-founded or transitive irreflexive relations don't have cycles (#18512) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Since `simp` can prove statements such as `cycle r [a, b, c] ↔ r a b ∧ r b c ∧ r c a`, this gives us a nice generalization of asymmetry. Co-authored-by: Eric Wieser --- src/data/list/cycle.lean | 24 ++++++++++++++++++++++++ src/order/rel_classes.lean | 6 +++++- 2 files changed, 29 insertions(+), 1 deletion(-) diff --git a/src/data/list/cycle.lean b/src/data/list/cycle.lean index 173833768ea84..80d88fab41dfc 100644 --- a/src/data/list/cycle.lean +++ b/src/data/list/cycle.lean @@ -805,6 +805,19 @@ by rw [range_succ, ←coe_cons_eq_coe_append, chain_coe_cons, ←range_succ, cha variables {r : α → α → Prop} {s : cycle α} +theorem chain.imp {r₁ r₂ : α → α → Prop} (H : ∀ a b, r₁ a b → r₂ a b) (p : chain r₁ s) : + chain r₂ s := +begin + induction s using cycle.induction_on, + { triv }, + { rw chain_coe_cons at ⊢ p, + exact p.imp H } +end + +/-- As a function from a relation to a predicate, `chain` is monotonic. -/ +theorem chain_mono : monotone (chain : (α → α → Prop) → cycle α → Prop) := +λ a b hab s, chain.imp hab + theorem chain_of_pairwise : (∀ (a ∈ s) (b ∈ s), r a b) → chain r s := begin induction s using cycle.induction_on with a l _, @@ -845,6 +858,17 @@ theorem chain_iff_pairwise [is_trans α r] : chain r s ↔ ∀ (a ∈ s) (b ∈ { exact trans (hs.2.2 b hb) (hs.1 c (or.inl hc)) } end, cycle.chain_of_pairwise⟩ +theorem chain.eq_nil_of_irrefl [is_trans α r] [is_irrefl α r] (h : chain r s) : s = nil := +begin + induction s using cycle.induction_on with a l _ h, + { refl }, + { have ha := mem_cons_self a l, + exact (irrefl_of r a $ chain_iff_pairwise.1 h a ha a ha).elim } +end + +theorem chain.eq_nil_of_well_founded [is_well_founded α r] (h : chain r s) : s = nil := +chain.eq_nil_of_irrefl $ h.imp $ λ _ _, relation.trans_gen.single + theorem forall_eq_of_chain [is_trans α r] [is_antisymm α r] (hs : chain r s) {a b : α} (ha : a ∈ s) (hb : b ∈ s) : a = b := by { rw chain_iff_pairwise at hs, exact antisymm (hs a ha b hb) (hs b hb a ha) } diff --git a/src/order/rel_classes.lean b/src/order/rel_classes.lean index b22650778f893..263bf5e1487cf 100644 --- a/src/order/rel_classes.lean +++ b/src/order/rel_classes.lean @@ -3,8 +3,9 @@ Copyright (c) 2020 Jeremy Avigad. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Mario Carneiro, Yury G. Kudryashov -/ -import order.basic import logic.is_empty +import logic.relation +import order.basic /-! # Unbundled relation classes @@ -247,6 +248,9 @@ instance is_well_founded.is_asymm (r : α → α → Prop) [is_well_founded α r instance is_well_founded.is_irrefl (r : α → α → Prop) [is_well_founded α r] : is_irrefl α r := is_asymm.is_irrefl +instance (r : α → α → Prop) [i : is_well_founded α r] : is_well_founded α (relation.trans_gen r) := +⟨i.wf.trans_gen⟩ + /-- A class for a well founded relation `<`. -/ @[reducible] def well_founded_lt (α : Type*) [has_lt α] : Prop := is_well_founded α (<) From ec4528061e02f0acc848ed06eb22573645602c7e Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Wed, 3 May 2023 03:14:46 +0000 Subject: [PATCH 0931/1291] feat(analysis/special_functions): file for evaluations of specific improper integrals (#18821) We have the file `analysis/special_functions/integrals.lean` which has numerous evaluations of interval integrals. This adds a counterpart file for integrals over infinite intervals. (Added as preparation for a forthcoming PR about Mellin transforms -- I wanted somewhere to put the fact that `x ^ (-a)` is integrable on `[1, infinity)`). Co-authored-by: sgouezel Co-authored-by: Eric Wieser --- src/analysis/special_functions/gamma.lean | 10 +- .../special_functions/improper_integrals.lean | 111 ++++++++++++++++++ src/measure_theory/measure/lebesgue.lean | 27 +++++ 3 files changed, 140 insertions(+), 8 deletions(-) create mode 100644 src/analysis/special_functions/improper_integrals.lean diff --git a/src/analysis/special_functions/gamma.lean b/src/analysis/special_functions/gamma.lean index cecff8716528e..b47ada18e93b6 100644 --- a/src/analysis/special_functions/gamma.lean +++ b/src/analysis/special_functions/gamma.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: David Loeffler -/ import measure_theory.integral.exp_decay +import analysis.special_functions.improper_integrals import analysis.convolution import analysis.special_functions.trigonometric.euler_sine_prod @@ -61,13 +62,6 @@ noncomputable theory open filter interval_integral set real measure_theory asymptotics open_locale nat topology ennreal big_operators complex_conjugate -lemma integral_exp_neg_Ioi : ∫ (x : ℝ) in Ioi 0, exp (-x) = 1 := -begin - refine tendsto_nhds_unique (interval_integral_tendsto_integral_Ioi _ _ tendsto_id) _, - { simpa only [neg_mul, one_mul] using exp_neg_integrable_on_Ioi 0 zero_lt_one, }, - { simpa using tendsto_exp_neg_at_top_nhds_0.const_sub 1, }, -end - namespace real /-- Asymptotic bound for the `Γ` function integrand. -/ @@ -161,7 +155,7 @@ end lemma Gamma_integral_one : Gamma_integral 1 = 1 := by simpa only [←of_real_one, Gamma_integral_of_real, of_real_inj, sub_self, - rpow_zero, mul_one] using integral_exp_neg_Ioi + rpow_zero, mul_one] using integral_exp_neg_Ioi_zero end complex diff --git a/src/analysis/special_functions/improper_integrals.lean b/src/analysis/special_functions/improper_integrals.lean new file mode 100644 index 0000000000000..1d10e1374d9ea --- /dev/null +++ b/src/analysis/special_functions/improper_integrals.lean @@ -0,0 +1,111 @@ +/- +Copyright (c) 2023 David Loeffler. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: David Loeffler +-/ +import measure_theory.integral.integral_eq_improper +import measure_theory.group.integration +import measure_theory.integral.exp_decay +import analysis.special_functions.integrals + +/-! +# Evaluation of specific improper integrals + +This file contains some integrability results, and evaluations of integrals, over `ℝ` or over +half-infinite intervals in `ℝ`. + +## See also + +- `analysis.special_functions.integrals` -- integrals over finite intervals +- `analysis.special_functions.gaussian` -- integral of `exp (-x ^ 2)` +- `analysis.special_functions.japanese_bracket`-- integrability of `(1+‖x‖)^(-r)`. +-/ + +open real set filter measure_theory interval_integral +open_locale topology + +lemma integrable_on_exp_Iic (c : ℝ) : integrable_on exp (Iic c) := +begin + refine integrable_on_Iic_of_interval_integral_norm_bounded (exp c) c (λ y, + interval_integrable_exp.1) tendsto_id (eventually_of_mem (Iic_mem_at_bot 0) (λ y hy, _)), + simp_rw [(norm_of_nonneg (exp_pos _).le), integral_exp, sub_le_self_iff], + exact (exp_pos _).le, +end + +lemma integral_exp_Iic (c : ℝ) : ∫ (x : ℝ) in Iic c, exp x = exp c := +begin + refine tendsto_nhds_unique (interval_integral_tendsto_integral_Iic _ (integrable_on_exp_Iic _) + tendsto_id) _, + simp_rw [integral_exp, (show 𝓝 (exp c) = 𝓝 (exp c - 0), by rw sub_zero)], + exact tendsto_exp_at_bot.const_sub _, +end + +lemma integral_exp_Iic_zero : ∫ (x : ℝ) in Iic 0, exp x = 1 := exp_zero ▸ integral_exp_Iic 0 + +lemma integral_exp_neg_Ioi (c : ℝ) : ∫ (x : ℝ) in Ioi c, exp (-x) = exp (-c) := +by simpa only [integral_comp_neg_Ioi] using integral_exp_Iic (-c) + +lemma integral_exp_neg_Ioi_zero : ∫ (x : ℝ) in Ioi 0, exp (-x) = 1 := +by simpa only [neg_zero, exp_zero] using integral_exp_neg_Ioi 0 + +/-- If `0 < c`, then `(λ t : ℝ, t ^ a)` is integrable on `(c, ∞)` for all `a < -1`. -/ +lemma integrable_on_Ioi_rpow_of_lt {a : ℝ} (ha : a < -1) {c : ℝ} (hc : 0 < c) : + integrable_on (λ t : ℝ, t ^ a) (Ioi c) := +begin + have hd : ∀ (x : ℝ) (hx : x ∈ Ici c), has_deriv_at (λ t, t ^ (a + 1) / (a + 1)) (x ^ a) x, + { intros x hx, + convert (has_deriv_at_rpow_const (or.inl (hc.trans_le hx).ne')).div_const _, + field_simp [show a + 1 ≠ 0, from ne_of_lt (by linarith), mul_comm] }, + have ht : tendsto (λ t, t ^ (a + 1) / (a + 1)) at_top (𝓝 (0/(a+1))), + { apply tendsto.div_const, + simpa only [neg_neg] using tendsto_rpow_neg_at_top (by linarith : 0 < -(a + 1)) }, + exact integrable_on_Ioi_deriv_of_nonneg' hd (λ t ht, rpow_nonneg_of_nonneg (hc.trans ht).le a) ht +end + +lemma integral_Ioi_rpow_of_lt {a : ℝ} (ha : a < -1) {c : ℝ} (hc : 0 < c) : + ∫ (t : ℝ) in Ioi c, t ^ a = -c ^ (a + 1) / (a + 1) := +begin + have hd : ∀ (x : ℝ) (hx : x ∈ Ici c), has_deriv_at (λ t, t ^ (a + 1) / (a + 1)) (x ^ a) x, + { intros x hx, + convert (has_deriv_at_rpow_const (or.inl (hc.trans_le hx).ne')).div_const _, + field_simp [show a + 1 ≠ 0, from ne_of_lt (by linarith), mul_comm] }, + have ht : tendsto (λ t, t ^ (a + 1) / (a + 1)) at_top (𝓝 (0/(a+1))), + { apply tendsto.div_const, + simpa only [neg_neg] using tendsto_rpow_neg_at_top (by linarith : 0 < -(a + 1)) }, + convert integral_Ioi_of_has_deriv_at_of_tendsto' hd (integrable_on_Ioi_rpow_of_lt ha hc) ht, + simp only [neg_div, zero_div, zero_sub], +end + +lemma integrable_on_Ioi_cpow_of_lt {a : ℂ} (ha : a.re < -1) {c : ℝ} (hc : 0 < c) : + integrable_on (λ t : ℝ, (t : ℂ) ^ a) (Ioi c) := +begin + rw [integrable_on, ←integrable_norm_iff, ←integrable_on], + refine (integrable_on_Ioi_rpow_of_lt ha hc).congr_fun (λ x hx, _) measurable_set_Ioi, + { dsimp only, + rw [complex.norm_eq_abs, complex.abs_cpow_eq_rpow_re_of_pos (hc.trans hx)] }, + { refine continuous_on.ae_strongly_measurable (λ t ht, _) measurable_set_Ioi, + exact (complex.continuous_at_of_real_cpow_const _ _ + (or.inr (hc.trans ht).ne')).continuous_within_at } +end + +lemma integral_Ioi_cpow_of_lt {a : ℂ} (ha : a.re < -1) {c : ℝ} (hc : 0 < c) : + ∫ (t : ℝ) in Ioi c, (t : ℂ) ^ a = -(c : ℂ) ^ (a + 1) / (a + 1) := +begin + refine tendsto_nhds_unique (interval_integral_tendsto_integral_Ioi c + (integrable_on_Ioi_cpow_of_lt ha hc) tendsto_id) _, + suffices : tendsto (λ (x : ℝ), ((x : ℂ) ^ (a + 1) - (c : ℂ) ^ (a + 1)) / (a + 1)) at_top + (𝓝 $ -c ^ (a + 1) / (a + 1)), + { refine this.congr' ((eventually_gt_at_top 0).mp (eventually_of_forall $ λ x hx, _)), + rw [integral_cpow, id.def], + refine or.inr ⟨_, not_mem_uIcc_of_lt hc hx⟩, + apply_fun complex.re, + rw [complex.neg_re, complex.one_re], + exact ha.ne }, + simp_rw [←zero_sub, sub_div], + refine (tendsto.div_const _ _).sub_const _, + rw tendsto_zero_iff_norm_tendsto_zero, + refine (tendsto_rpow_neg_at_top (by linarith : 0 < -(a.re + 1))).congr' + ((eventually_gt_at_top 0).mp (eventually_of_forall $ λ x hx, _)), + simp_rw [neg_neg, complex.norm_eq_abs, complex.abs_cpow_eq_rpow_re_of_pos hx, + complex.add_re, complex.one_re], +end diff --git a/src/measure_theory/measure/lebesgue.lean b/src/measure_theory/measure/lebesgue.lean index e9cd6f6834cd9..c347ccce0c02b 100644 --- a/src/measure_theory/measure/lebesgue.lean +++ b/src/measure_theory/measure/lebesgue.lean @@ -630,3 +630,30 @@ begin end end summable_norm_Icc + +/-! +### Substituting `-x` for `x` + +These lemmas are stated in terms of either `Iic` or `Ioi` (neglecting `Iio` and `Ici`) to match +mathlib's conventions for integrals over finite intervals (see `interval_integral`). For the case +of finite integrals, see `interval_integral.integral_comp_neg`. +-/ + +@[simp] lemma integral_comp_neg_Iic {E : Type*} + [normed_add_comm_group E] [normed_space ℝ E] [complete_space E] (c : ℝ) (f : ℝ → E) : + ∫ x in Iic c, f (-x) = ∫ x in Ioi (-c), f x := +begin + have A : measurable_embedding (λ x : ℝ, -x), + from (homeomorph.neg ℝ).closed_embedding.measurable_embedding, + have := A.set_integral_map f (Ici (-c)), + rw measure.map_neg_eq_self (volume : measure ℝ) at this, + simp_rw [←integral_Ici_eq_integral_Ioi, this, neg_preimage, preimage_neg_Ici, neg_neg], +end + +@[simp] lemma integral_comp_neg_Ioi {E : Type*} + [normed_add_comm_group E] [normed_space ℝ E] [complete_space E] (c : ℝ) (f : ℝ → E) : + ∫ x in Ioi c, f (-x) = ∫ x in Iic (-c), f x := +begin + rw [←neg_neg c, ←integral_comp_neg_Iic], + simp only [neg_neg], +end From f69db8cecc668e2d5894d7e9bfc491da60db3b9f Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Hagb=20=28Junyu=20Guo=20=E9=83=AD=E4=BF=8A=E4=BD=99=29?= Date: Wed, 3 May 2023 03:14:47 +0000 Subject: [PATCH 0932/1291] feat(data/mv_polynomial/basic): add and generalize some lemmas from finsupp and monoid_algebra (#18855) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Most of these changes generalize from `distrib_mul_action` to `smul_zero_class`. The new lemmas are all just proved using corresponding lemmas on the underlying types. Co-authored-by: Hagb (Junyu Guo 郭俊余) --- src/algebra/monoid_algebra/basic.lean | 8 +++---- src/data/finsupp/basic.lean | 2 +- src/data/mv_polynomial/basic.lean | 32 +++++++++++++++++++-------- 3 files changed, 28 insertions(+), 14 deletions(-) diff --git a/src/algebra/monoid_algebra/basic.lean b/src/algebra/monoid_algebra/basic.lean index 53dce47b0586b..a23033c2a43ba 100644 --- a/src/algebra/monoid_algebra/basic.lean +++ b/src/algebra/monoid_algebra/basic.lean @@ -1130,24 +1130,24 @@ instance [monoid R] [semiring k] [distrib_mul_action R k] : distrib_mul_action R (add_monoid_algebra k G) := finsupp.distrib_mul_action G k -instance [monoid R] [semiring k] [distrib_mul_action R k] [has_faithful_smul R k] [nonempty G] : +instance [semiring k] [smul_zero_class R k] [has_faithful_smul R k] [nonempty G] : has_faithful_smul R (add_monoid_algebra k G) := finsupp.has_faithful_smul instance [semiring R] [semiring k] [module R k] : module R (add_monoid_algebra k G) := finsupp.module G k -instance [monoid R] [monoid S] [semiring k] [distrib_mul_action R k] [distrib_mul_action S k] +instance [semiring k] [smul_zero_class R k] [smul_zero_class S k] [has_smul R S] [is_scalar_tower R S k] : is_scalar_tower R S (add_monoid_algebra k G) := finsupp.is_scalar_tower G k -instance [monoid R] [monoid S] [semiring k] [distrib_mul_action R k] [distrib_mul_action S k] +instance [semiring k] [smul_zero_class R k] [smul_zero_class S k] [smul_comm_class R S k] : smul_comm_class R S (add_monoid_algebra k G) := finsupp.smul_comm_class G k -instance [monoid R] [semiring k] [distrib_mul_action R k] [distrib_mul_action Rᵐᵒᵖ k] +instance [semiring k] [smul_zero_class R k] [smul_zero_class Rᵐᵒᵖ k] [is_central_scalar R k] : is_central_scalar R (add_monoid_algebra k G) := finsupp.is_central_scalar G k diff --git a/src/data/finsupp/basic.lean b/src/data/finsupp/basic.lean index 7d90e3d52214d..8e49f3d64eff3 100644 --- a/src/data/finsupp/basic.lean +++ b/src/data/finsupp/basic.lean @@ -1306,7 +1306,7 @@ instance [semiring R] [add_comm_monoid M] [module R M] : module R (α →₀ M) variables {α M} {R} -lemma support_smul {_ : monoid R} [add_monoid M] [distrib_mul_action R M] {b : R} {g : α →₀ M} : +lemma support_smul [add_monoid M] [smul_zero_class R M] {b : R} {g : α →₀ M} : (b • g).support ⊆ g.support := λ a, by { simp only [smul_apply, mem_support_iff, ne.def], exact mt (λ h, h.symm ▸ smul_zero _) } diff --git a/src/data/mv_polynomial/basic.lean b/src/data/mv_polynomial/basic.lean index d34f0b96e9fc0..1be5b486ae32f 100644 --- a/src/data/mv_polynomial/basic.lean +++ b/src/data/mv_polynomial/basic.lean @@ -102,24 +102,27 @@ instance [monoid R] [comm_semiring S₁] [distrib_mul_action R S₁] : distrib_mul_action R (mv_polynomial σ S₁) := add_monoid_algebra.distrib_mul_action -instance [monoid R] [comm_semiring S₁] [distrib_mul_action R S₁] [has_faithful_smul R S₁] : +instance [comm_semiring S₁] [smul_zero_class R S₁] : smul_zero_class R (mv_polynomial σ S₁) := +add_monoid_algebra.smul_zero_class + +instance [comm_semiring S₁] [smul_zero_class R S₁] [has_faithful_smul R S₁] : has_faithful_smul R (mv_polynomial σ S₁) := add_monoid_algebra.has_faithful_smul instance [semiring R] [comm_semiring S₁] [module R S₁] : module R (mv_polynomial σ S₁) := add_monoid_algebra.module -instance [monoid R] [monoid S₁] [comm_semiring S₂] - [has_smul R S₁] [distrib_mul_action R S₂] [distrib_mul_action S₁ S₂] [is_scalar_tower R S₁ S₂] : +instance [comm_semiring S₂] + [has_smul R S₁] [smul_zero_class R S₂] [smul_zero_class S₁ S₂] [is_scalar_tower R S₁ S₂] : is_scalar_tower R S₁ (mv_polynomial σ S₂) := add_monoid_algebra.is_scalar_tower -instance [monoid R] [monoid S₁][comm_semiring S₂] - [distrib_mul_action R S₂] [distrib_mul_action S₁ S₂] [smul_comm_class R S₁ S₂] : +instance [comm_semiring S₂] + [smul_zero_class R S₂] [smul_zero_class S₁ S₂] [smul_comm_class R S₁ S₂] : smul_comm_class R S₁ (mv_polynomial σ S₂) := add_monoid_algebra.smul_comm_class -instance [monoid R] [comm_semiring S₁] [distrib_mul_action R S₁] [distrib_mul_action Rᵐᵒᵖ S₁] +instance [comm_semiring S₁] [smul_zero_class R S₁] [smul_zero_class Rᵐᵒᵖ S₁] [is_central_scalar R S₁] : is_central_scalar R (mv_polynomial σ S₁) := add_monoid_algebra.is_central_scalar @@ -218,6 +221,9 @@ lemma smul_eq_C_mul (p : mv_polynomial σ R) (a : R) : a • p = C a * p := C_mu lemma C_eq_smul_one : (C a : mv_polynomial σ R) = a • 1 := by rw [← C_mul', mul_one] +lemma smul_monomial {S₁ : Type*} [smul_zero_class S₁ R] (r : S₁) : + r • monomial s a = monomial s (r • a) := finsupp.smul_single _ _ _ + lemma X_injective [nontrivial R] : function.injective (X : σ → mv_polynomial σ R) := (monomial_left_injective one_ne_zero).comp (finsupp.single_left_injective one_ne_zero) @@ -422,7 +428,7 @@ by rw [X_pow_eq_monomial, support_monomial, if_neg (one_ne_zero' R)] @[simp] lemma support_zero : (0 : mv_polynomial σ R).support = ∅ := rfl -lemma support_smul [distrib_mul_action R S₁] {a : R} {f : mv_polynomial σ S₁} : +lemma support_smul {S₁ : Type*} [smul_zero_class S₁ R] {a : S₁} {f : mv_polynomial σ R} : (a • f).support ⊆ f.support := finsupp.support_smul @@ -463,7 +469,7 @@ lemma ext_iff (p q : mv_polynomial σ R) : @[simp] lemma coeff_add (m : σ →₀ ℕ) (p q : mv_polynomial σ R) : coeff m (p + q) = coeff m p + coeff m q := add_apply p q m -@[simp] lemma coeff_smul {S₁ : Type*} [monoid S₁] [distrib_mul_action S₁ R] +@[simp] lemma coeff_smul {S₁ : Type*} [smul_zero_class S₁ R] (m : σ →₀ ℕ) (c : S₁) (p : mv_polynomial σ R) : coeff m (c • p) = c • coeff m p := smul_apply c p m @@ -557,6 +563,10 @@ add_monoid_algebra.support_mul_single p _ (by simp) _ (X s * p).support = p.support.map (add_left_embedding (single s 1)) := add_monoid_algebra.support_single_mul p _ (by simp) _ +@[simp] lemma support_smul_eq {S₁ : Type*} [semiring S₁] [module S₁ R] [no_zero_smul_divisors S₁ R] + {a : S₁} (h : a ≠ 0) (p : mv_polynomial σ R) : (a • p).support = p.support := +finsupp.support_smul_eq h + lemma support_sdiff_support_subset_support_add [decidable_eq σ] (p q : mv_polynomial σ R) : p.support \ q.support ⊆ (p + q).support := begin @@ -626,6 +636,9 @@ lemma ne_zero_iff {p : mv_polynomial σ R} : p ≠ 0 ↔ ∃ d, coeff d p ≠ 0 := by { rw [ne.def, eq_zero_iff], push_neg, } +@[simp] lemma support_eq_empty {p : mv_polynomial σ R} : p.support = ∅ ↔ p = 0 := +finsupp.support_eq_empty + lemma exists_coeff_ne_zero {p : mv_polynomial σ R} (h : p ≠ 0) : ∃ d, coeff d p ≠ 0 := ne_zero_iff.mp h @@ -675,7 +688,8 @@ variables (R) by simp [constant_coeff_eq] variables {R} -@[simp] lemma constant_coeff_smul [distrib_mul_action R S₁] (a : R) (f : mv_polynomial σ S₁) : +@[simp] lemma constant_coeff_smul {R : Type*} [smul_zero_class R S₁] + (a : R) (f : mv_polynomial σ S₁) : constant_coeff (a • f) = a • constant_coeff f := rfl lemma constant_coeff_monomial [decidable_eq σ] (d : σ →₀ ℕ) (r : R) : From 645b6de46f671ac56ffea84c8752ee1cc2b8c898 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 3 May 2023 03:14:48 +0000 Subject: [PATCH 0933/1291] refactor(algebra/order/to_interval_mod): Negate the definition of `mem_Ioo_mod` (#18912) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This replaces `mem_Ioo_mod hp a b` with `¬add_comm_group.modeq p a b`. This is more consistent with `int.modeq`, `nat.modeq`, and `smodeq`. There's still some duplication here, but at least these four ideas are now conceptually aligned. This remove any lemmas of the form `¬modeq p a b ↔ _ ≠ _` as these are now trivial consequences of the `modeq p a b ↔ _ = _` versions, --- src/algebra/order/to_interval_mod.lean | 131 +++++++++++++------------ src/topology/instances/add_circle.lean | 11 ++- 2 files changed, 74 insertions(+), 68 deletions(-) diff --git a/src/algebra/order/to_interval_mod.lean b/src/algebra/order/to_interval_mod.lean index 201deaa40340c..4943b56adce75 100644 --- a/src/algebra/order/to_interval_mod.lean +++ b/src/algebra/order/to_interval_mod.lean @@ -27,7 +27,12 @@ interval. * `to_Ioc_div hp a b` (where `hp : 0 < p`): The unique integer such that this multiple of `p`, subtracted from `b`, is in `Ioc a (a + p)`. * `to_Ioc_mod hp a b` (where `hp : 0 < p`): Reduce `b` to the interval `Ioc a (a + p)`. +* `add_comm_group.modeq p a b`: `a` and `b` are congruent modulo a multiple of `p`. See also + `smodeq` which is a more general version in arbitrary submodules. +## TODO + +Unify `smodeq` and `add_comm_group.modeq`, which were originally developed independently. -/ noncomputable theory @@ -338,118 +343,118 @@ end /-! ### Links between the `Ico` and `Ioc` variants applied to the same element -/ section Ico_Ioc -variables (a b) +namespace add_comm_group +variables (a b) omit hα -/-- `mem_Ioo_mod p a b` means that `b` lies in the open interval `(a, a + p)` modulo `p`. -Equivalently (as shown below), `b` is not congruent to `a` modulo `p`, or `to_Ico_mod hp a` agrees -with `to_Ioc_mod hp a` at `b`, or `to_Ico_div hp a` agrees with `to_Ioc_div hp a` at `b`. -/ -def mem_Ioo_mod (p a b : α) : Prop := ∃ z : ℤ, b - z • p ∈ set.Ioo a (a + p) +/-- `add_comm_group.modeq p a b` means that `b` does not lie in the open interval `(a, a + p)` +modulo `p`. + +Equivalently (as shown below), `b` is congruent to `a` modulo `p`, or `to_Ico_mod hp a` disagrees +with `to_Ioc_mod hp a` at `b`, or `to_Ico_div hp a` disagrees with `to_Ioc_div hp a` at `b`. -/ +def modeq (p a b : α) : Prop := ∀ z : ℤ, b - z • p ∉ set.Ioo a (a + p) include hα -lemma tfae_mem_Ioo_mod : - tfae [mem_Ioo_mod p a b, - to_Ico_mod hp a b = to_Ioc_mod hp a b, - to_Ico_mod hp a b + p ≠ to_Ioc_mod hp a b, - to_Ico_mod hp a b ≠ a] := +lemma tfae_modeq : + tfae [modeq p a b, + to_Ico_mod hp a b ≠ to_Ioc_mod hp a b, + to_Ico_mod hp a b + p = to_Ioc_mod hp a b, + to_Ico_mod hp a b = a] := begin - tfae_have : 1 → 2, - { exact λ ⟨i, hi⟩, + rw modeq, + tfae_have : 2 → 1, + { rw [←not_exists, not_imp_not], + exact λ ⟨i, hi⟩, ((to_Ico_mod_eq_iff hp).2 ⟨set.Ioo_subset_Ico_self hi, i, (sub_add_cancel b _).symm⟩).trans ((to_Ioc_mod_eq_iff hp).2 ⟨set.Ioo_subset_Ioc_self hi, i, (sub_add_cancel b _).symm⟩).symm }, - tfae_have : 2 → 3, - { intro h, rw [h, ne, add_right_eq_self], exact hp.ne' }, - tfae_have : 3 → 4, - { refine mt (λ h, _), + tfae_have : 3 → 2, + { intro h, rw [←h, ne, eq_comm, add_right_eq_self], exact hp.ne' }, + tfae_have : 4 → 3, + { intro h, rw [h, eq_comm, to_Ioc_mod_eq_iff, set.right_mem_Ioc], refine ⟨lt_add_of_pos_right a hp, to_Ico_div hp a b - 1, _⟩, rw [sub_one_zsmul, add_add_add_comm, add_right_neg, add_zero], conv_lhs { rw [← to_Ico_mod_add_to_Ico_div_zsmul hp a b, h] } }, - tfae_have : 4 → 1, - { have h' := to_Ico_mod_mem_Ico hp a b, exact λ h, ⟨_, h'.1.lt_of_ne' h, h'.2⟩ }, + tfae_have : 1 → 4, + { rw [←not_exists, not_imp_comm], + have h' := to_Ico_mod_mem_Ico hp a b, + exact λ h, ⟨_, h'.1.lt_of_ne' h, h'.2⟩ }, tfae_finish, end variables {a b} -lemma mem_Ioo_mod_iff_to_Ico_mod_eq_to_Ioc_mod : - mem_Ioo_mod p a b ↔ to_Ico_mod hp a b = to_Ioc_mod hp a b := (tfae_mem_Ioo_mod hp a b).out 0 1 -lemma mem_Ioo_mod_iff_to_Ico_mod_add_period_ne_to_Ioc_mod : - mem_Ioo_mod p a b ↔ to_Ico_mod hp a b + p ≠ to_Ioc_mod hp a b := (tfae_mem_Ioo_mod hp a b).out 0 2 -lemma mem_Ioo_mod_iff_to_Ico_mod_ne_left : - mem_Ioo_mod p a b ↔ to_Ico_mod hp a b ≠ a := (tfae_mem_Ioo_mod hp a b).out 0 3 +lemma modeq_iff_to_Ico_mod_ne_to_Ioc_mod : + modeq p a b ↔ to_Ico_mod hp a b ≠ to_Ioc_mod hp a b := (tfae_modeq hp a b).out 0 1 +lemma modeq_iff_to_Ico_mod_add_period_eq_to_Ioc_mod : + modeq p a b ↔ to_Ico_mod hp a b + p = to_Ioc_mod hp a b := (tfae_modeq hp a b).out 0 2 +lemma modeq_iff_to_Ico_mod_eq_left : + modeq p a b ↔ to_Ico_mod hp a b = a := (tfae_modeq hp a b).out 0 3 -lemma not_mem_Ioo_mod_iff_to_Ico_mod_add_period_eq_to_Ioc_mod : - ¬mem_Ioo_mod p a b ↔ to_Ico_mod hp a b + p = to_Ioc_mod hp a b := -(mem_Ioo_mod_iff_to_Ico_mod_add_period_ne_to_Ioc_mod hp).not_left +lemma not_modeq_iff_to_Ico_mod_eq_to_Ioc_mod : + ¬modeq p a b ↔ to_Ico_mod hp a b = to_Ioc_mod hp a b := +(modeq_iff_to_Ico_mod_ne_to_Ioc_mod _).not_left -lemma not_mem_Ioo_mod_iff_to_Ico_mod_eq_left : ¬mem_Ioo_mod p a b ↔ to_Ico_mod hp a b = a := -(mem_Ioo_mod_iff_to_Ico_mod_ne_left hp).not_left - -lemma mem_Ioo_mod_iff_to_Ioc_mod_ne_right : mem_Ioo_mod p a b ↔ to_Ioc_mod hp a b ≠ a + p := +lemma modeq_iff_to_Ioc_mod_eq_right : modeq p a b ↔ to_Ioc_mod hp a b = a + p := begin - rw [mem_Ioo_mod_iff_to_Ico_mod_eq_to_Ioc_mod, to_Ico_mod_eq_iff hp], + rw [modeq_iff_to_Ico_mod_ne_to_Ioc_mod hp, ne, to_Ico_mod_eq_iff hp, not_iff_comm], obtain ⟨h₁, h₂⟩ := to_Ioc_mod_mem_Ioc hp a b, - exact ⟨λ h, h.1.2.ne, λ h, ⟨⟨h₁.le, h₂.lt_of_ne h⟩, _, - (to_Ioc_mod_add_to_Ioc_div_zsmul _ _ _).symm⟩⟩, + exact ⟨λ h, ⟨⟨h₁.le, h₂.lt_of_ne h⟩, _, (to_Ioc_mod_add_to_Ioc_div_zsmul _ _ _).symm⟩, + λ h, h.1.2.ne⟩, end -lemma not_mem_Ioo_mod_iff_to_Ioc_eq_right : ¬mem_Ioo_mod p a b ↔ to_Ioc_mod hp a b = a + p := -(mem_Ioo_mod_iff_to_Ioc_mod_ne_right hp).not_left - -lemma mem_Ioo_mod_iff_to_Ico_div_eq_to_Ioc_div : - mem_Ioo_mod p a b ↔ to_Ico_div hp a b = to_Ioc_div hp a b := -by rw [mem_Ioo_mod_iff_to_Ico_mod_eq_to_Ioc_mod hp, +lemma not_modeq_iff_to_Ico_div_eq_to_Ioc_div : + ¬modeq p a b ↔ to_Ico_div hp a b = to_Ioc_div hp a b := +by rw [not_modeq_iff_to_Ico_mod_eq_to_Ioc_mod hp, to_Ico_mod, to_Ioc_mod, sub_right_inj, (zsmul_strict_mono_left hp).injective.eq_iff] -lemma mem_Ioo_mod_iff_to_Ico_div_ne_to_Ioc_div_add_one : - mem_Ioo_mod p a b ↔ to_Ico_div hp a b ≠ to_Ioc_div hp a b + 1 := -by rw [mem_Ioo_mod_iff_to_Ico_mod_add_period_ne_to_Ioc_mod hp, ne, ne, to_Ico_mod, to_Ioc_mod, +lemma modeq_iff_to_Ico_div_eq_to_Ioc_div_add_one : + modeq p a b ↔ to_Ico_div hp a b = to_Ioc_div hp a b + 1 := +by rw [modeq_iff_to_Ico_mod_add_period_eq_to_Ioc_mod hp, to_Ico_mod, to_Ioc_mod, ← eq_sub_iff_add_eq, sub_sub, sub_right_inj, ← add_one_zsmul, (zsmul_strict_mono_left hp).injective.eq_iff] -lemma not_mem_Ioo_mod_iff_to_Ico_div_eq_to_Ioc_div_add_one : - ¬mem_Ioo_mod p a b ↔ to_Ico_div hp a b = to_Ioc_div hp a b + 1 := -(mem_Ioo_mod_iff_to_Ico_div_ne_to_Ioc_div_add_one hp).not_left - include hp -lemma mem_Ioo_mod_iff_ne_add_zsmul : mem_Ioo_mod p a b ↔ ∀ z : ℤ, b ≠ a + z • p := +lemma modeq_iff_eq_add_zsmul : modeq p a b ↔ ∃ z : ℤ, b = a + z • p := begin - rw [mem_Ioo_mod_iff_to_Ico_mod_ne_left hp, ← not_iff_not], - push_neg, split; intro h, + rw [modeq_iff_to_Ico_mod_eq_left hp], + split; intro h, { rw ← h, exact ⟨_, (to_Ico_mod_add_to_Ico_div_zsmul _ _ _).symm⟩ }, { rw [to_Ico_mod_eq_iff, set.left_mem_Ico], exact ⟨lt_add_of_pos_right a hp, h⟩, }, end -lemma not_mem_Ioo_mod_iff_eq_add_zsmul : ¬mem_Ioo_mod p a b ↔ ∃ z : ℤ, b = a + z • p := -by simpa only [not_forall, not_ne_iff] using (mem_Ioo_mod_iff_ne_add_zsmul hp).not +lemma not_modeq_iff_ne_add_zsmul : ¬modeq p a b ↔ ∀ z : ℤ, b ≠ a + z • p := +by rw [modeq_iff_eq_add_zsmul hp, not_exists] -lemma not_mem_Ioo_mod_iff_eq_mod_zmultiples : - ¬mem_Ioo_mod p a b ↔ (b : α ⧸ add_subgroup.zmultiples p) = a := -by simp_rw [not_mem_Ioo_mod_iff_eq_add_zsmul hp, quotient_add_group.eq_iff_sub_mem, +lemma modeq_iff_eq_mod_zmultiples : + modeq p a b ↔ (b : α ⧸ add_subgroup.zmultiples p) = a := +by simp_rw [modeq_iff_eq_add_zsmul hp, quotient_add_group.eq_iff_sub_mem, add_subgroup.mem_zmultiples_iff, eq_sub_iff_add_eq', eq_comm] -lemma mem_Ioo_mod_iff_ne_mod_zmultiples : - mem_Ioo_mod p a b ↔ (b : α ⧸ add_subgroup.zmultiples p) ≠ a := -(not_mem_Ioo_mod_iff_eq_mod_zmultiples hp).not_right +lemma not_modeq_iff_ne_mod_zmultiples : + ¬modeq p a b ↔ (b : α ⧸ add_subgroup.zmultiples p) ≠ a := +(modeq_iff_eq_mod_zmultiples hp).not + +end add_comm_group lemma Ico_eq_locus_Ioc_eq_Union_Ioo : {b | to_Ico_mod hp a b = to_Ioc_mod hp a b} = ⋃ z : ℤ, set.Ioo (a + z • p) (a + p + z • p) := begin - ext1, simp_rw [set.mem_set_of, set.mem_Union, ← set.sub_mem_Ioo_iff_left], - exact (mem_Ioo_mod_iff_to_Ico_mod_eq_to_Ioc_mod hp).symm, + ext1, simp_rw [set.mem_set_of, set.mem_Union, ← set.sub_mem_Ioo_iff_left, + ←add_comm_group.not_modeq_iff_to_Ico_mod_eq_to_Ioc_mod, add_comm_group.modeq, not_forall, + not_not], end lemma to_Ioc_div_wcovby_to_Ico_div (a b : α) : to_Ioc_div hp a b ⩿ to_Ico_div hp a b := begin suffices : to_Ioc_div hp a b = to_Ico_div hp a b ∨ to_Ioc_div hp a b + 1 = to_Ico_div hp a b, { rwa [wcovby_iff_eq_or_covby, ←order.succ_eq_iff_covby] }, - rw [eq_comm, ←mem_Ioo_mod_iff_to_Ico_div_eq_to_Ioc_div, - eq_comm, ←not_mem_Ioo_mod_iff_to_Ico_div_eq_to_Ioc_div_add_one], - exact em _, + rw [eq_comm, ←add_comm_group.not_modeq_iff_to_Ico_div_eq_to_Ioc_div, + eq_comm, ←add_comm_group.modeq_iff_to_Ico_div_eq_to_Ioc_div_add_one], + exact em' _, end lemma to_Ico_mod_le_to_Ioc_mod (a b : α) : to_Ico_mod hp a b ≤ to_Ioc_mod hp a b := diff --git a/src/topology/instances/add_circle.lean b/src/topology/instances/add_circle.lean index 40a4bf35aa9d7..944ff5878b3b9 100644 --- a/src/topology/instances/add_circle.lean +++ b/src/topology/instances/add_circle.lean @@ -89,7 +89,8 @@ variables {x} (hx : (x : 𝕜 ⧸ zmultiples p) ≠ a) lemma to_Ico_mod_eventually_eq_to_Ioc_mod : to_Ico_mod hp a =ᶠ[𝓝 x] to_Ioc_mod hp a := is_open.mem_nhds (by {rw Ico_eq_locus_Ioc_eq_Union_Ioo, exact is_open_Union (λ i, is_open_Ioo)}) $ - (mem_Ioo_mod_iff_to_Ico_mod_eq_to_Ioc_mod hp).1 ((mem_Ioo_mod_iff_ne_mod_zmultiples hp).2 hx) + (add_comm_group.not_modeq_iff_to_Ico_mod_eq_to_Ioc_mod hp).1 $ + (add_comm_group.not_modeq_iff_ne_mod_zmultiples hp).2 hx lemma continuous_at_to_Ico_mod : continuous_at (to_Ico_mod hp a) x := let h := to_Ico_mod_eventually_eq_to_Ioc_mod hp a hx in continuous_at_iff_continuous_left_right.2 $ @@ -498,11 +499,11 @@ lemma equiv_Icc_quot_comp_mk_eq_to_Ioc_mod : equiv_Icc_quot p a ∘ quotient.mk' λ x, quot.mk _ ⟨to_Ioc_mod hp.out a x, Ioc_subset_Icc_self $ to_Ioc_mod_mem_Ioc _ _ x⟩ := begin rw equiv_Icc_quot_comp_mk_eq_to_Ico_mod, funext, - by_cases mem_Ioo_mod p a x, - { simp_rw (mem_Ioo_mod_iff_to_Ico_mod_eq_to_Ioc_mod hp.out).1 h }, - { simp_rw [not_imp_comm.1 (mem_Ioo_mod_iff_to_Ico_mod_ne_left hp.out).2 h, - not_imp_comm.1 (mem_Ioo_mod_iff_to_Ioc_mod_ne_right hp.out).2 h], + by_cases add_comm_group.modeq p a x, + { simp_rw [(add_comm_group.modeq_iff_to_Ico_mod_eq_left hp.out).1 h, + (add_comm_group.modeq_iff_to_Ioc_mod_eq_right hp.out).1 h], exact quot.sound endpoint_ident.mk }, + { simp_rw (add_comm_group.not_modeq_iff_to_Ico_mod_eq_to_Ioc_mod hp.out).1 h }, end /-- The natural map from `[a, a + p] ⊂ 𝕜` with endpoints identified to `𝕜 / ℤ • p`, as a From 61d8b8248633da198afea97ae7a90ee63bdf8c1c Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Wed, 3 May 2023 03:14:49 +0000 Subject: [PATCH 0934/1291] chore(ring_theory/polynomial/quotient): remove spurious restrict_scalars (#18916) Verifying that the `restrict_scalars` here is spurious. This proof is timing out badly in mathlib4. If someone would like to merge this, please do so, but please then handle updating the SHA in mathlib4 as well. :-) Co-authored-by: Scott Morrison --- src/ring_theory/polynomial/quotient.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/ring_theory/polynomial/quotient.lean b/src/ring_theory/polynomial/quotient.lean index ba8f0338e9692..f7cdff17c5fc2 100644 --- a/src/ring_theory/polynomial/quotient.lean +++ b/src/ring_theory/polynomial/quotient.lean @@ -22,7 +22,7 @@ variables {R : Type*} [comm_ring R] isomorphism of $R$-algebras $R[X] / \langle X - x \rangle \cong R$. -/ noncomputable def quotient_span_X_sub_C_alg_equiv (x : R) : (R[X] ⧸ ideal.span ({X - C x} : set R[X])) ≃ₐ[R] R := -(alg_equiv.restrict_scalars R $ ideal.quotient_equiv_alg_of_eq R +(ideal.quotient_equiv_alg_of_eq R (by exact ker_eval_ring_hom x : ring_hom.ker (aeval x).to_ring_hom = _)).symm.trans $ ideal.quotient_ker_alg_equiv_of_right_inverse $ λ _, eval_C From 15e02bcd7f0487553848c8fa2fe8f9db6b2c5284 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Wed, 3 May 2023 03:14:50 +0000 Subject: [PATCH 0935/1291] feat(probability/kernel/composition): add a notation for the product of kernels (#18918) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Define `κ ×ₖ η = probability_theory.kernel.prod κ η`. --- src/probability/kernel/composition.lean | 18 +++++++++++------- 1 file changed, 11 insertions(+), 7 deletions(-) diff --git a/src/probability/kernel/composition.lean b/src/probability/kernel/composition.lean index 719367cc72092..75b1c1930c9cb 100644 --- a/src/probability/kernel/composition.lean +++ b/src/probability/kernel/composition.lean @@ -15,7 +15,7 @@ We define * the map and comap of a kernel along a measurable function. * the composition `η ∘ₖ κ` of s-finite kernels `κ : kernel α β` and `η : kernel β γ`, a kernel from `α` to `γ`. -* the product `prod κ η` of s-finite kernels `κ : kernel α β` and `η : kernel α γ`, +* the product `κ ×ₖ η` of s-finite kernels `κ : kernel α β` and `η : kernel α γ`, a kernel from `α` to `β × γ`. A note on names: @@ -38,7 +38,7 @@ Kernels built from other kernels: We define a notation `η ∘ₖ κ = comp η κ`. `∫⁻ c, g c ∂((η ∘ₖ κ) a) = ∫⁻ b, ∫⁻ c, g c ∂(η b) ∂(κ a)` * `prod (κ : kernel α β) (η : kernel α γ) : kernel α (β × γ)`: product of 2 s-finite kernels. - `∫⁻ bc, f bc ∂(prod κ η a) = ∫⁻ b, ∫⁻ c, f (b, c) ∂(η a) ∂(κ a)` + `∫⁻ bc, f bc ∂((κ ×ₖ η) a) = ∫⁻ b, ∫⁻ c, f (b, c) ∂(η a) ∂(κ a)` ## Main statements @@ -53,6 +53,7 @@ Kernels built from other kernels: * `κ ⊗ₖ η = probability_theory.kernel.comp_prod κ η` * `η ∘ₖ κ = probability_theory.kernel.comp η κ` +* `κ ×ₖ η = probability_theory.kernel.prod κ η` -/ @@ -692,31 +693,34 @@ def prod (κ : kernel α β) [is_s_finite_kernel κ] (η : kernel α γ) [is_s_f kernel α (β × γ) := κ ⊗ₖ (swap_left (prod_mk_left η β)) +localized "infix (name := kernel.prod) ` ×ₖ `:100 := probability_theory.kernel.prod" in + probability_theory + lemma prod_apply (κ : kernel α β) [is_s_finite_kernel κ] (η : kernel α γ) [is_s_finite_kernel η] (a : α) {s : set (β × γ)} (hs : measurable_set s) : - (prod κ η) a s = ∫⁻ (b : β), (η a) {c : γ | (b, c) ∈ s} ∂(κ a) := + (κ ×ₖ η) a s = ∫⁻ (b : β), (η a) {c : γ | (b, c) ∈ s} ∂(κ a) := by simp_rw [prod, comp_prod_apply _ _ _ hs, swap_left_apply _ _, prod_mk_left_apply, prod.swap_prod_mk] lemma lintegral_prod (κ : kernel α β) [is_s_finite_kernel κ] (η : kernel α γ) [is_s_finite_kernel η] (a : α) {g : (β × γ) → ℝ≥0∞} (hg : measurable g) : - ∫⁻ c, g c ∂((prod κ η) a) = ∫⁻ b, ∫⁻ c, g (b, c) ∂(η a) ∂(κ a) := + ∫⁻ c, g c ∂((κ ×ₖ η) a) = ∫⁻ b, ∫⁻ c, g (b, c) ∂(η a) ∂(κ a) := by simp_rw [prod, lintegral_comp_prod _ _ _ hg, swap_left_apply, prod_mk_left_apply, prod.swap_prod_mk] instance is_markov_kernel.prod (κ : kernel α β) [is_markov_kernel κ] (η : kernel α γ) [is_markov_kernel η] : - is_markov_kernel (prod κ η) := + is_markov_kernel (κ ×ₖ η) := by { rw prod, apply_instance, } instance is_finite_kernel.prod (κ : kernel α β) [is_finite_kernel κ] (η : kernel α γ) [is_finite_kernel η] : - is_finite_kernel (prod κ η) := + is_finite_kernel (κ ×ₖ η) := by { rw prod, apply_instance, } instance is_s_finite_kernel.prod (κ : kernel α β) [is_s_finite_kernel κ] (η : kernel α γ) [is_s_finite_kernel η] : - is_s_finite_kernel (prod κ η) := + is_s_finite_kernel (κ ×ₖ η) := by { rw prod, apply_instance, } end prod From caa58cbf5bfb7f81ccbaca4e8b8ac4bc2b39cc1c Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 3 May 2023 03:14:51 +0000 Subject: [PATCH 0936/1291] chore(data/{complex,is_R_or_C}/basic): fix name of `eq_conj_iff_*` lemmas (#18922) These were all about `conj x = x` not `x = conj x`. --- src/analysis/complex/upper_half_plane/metric.lean | 2 +- src/analysis/inner_product_space/positive.lean | 2 +- src/analysis/inner_product_space/spectrum.lean | 2 +- src/analysis/inner_product_space/symmetric.lean | 4 ++-- src/analysis/special_functions/gamma.lean | 2 +- src/data/complex/basic.lean | 8 ++++---- src/data/complex/exponential.lean | 14 +++++++------- src/data/is_R_or_C/basic.lean | 6 +++--- src/field_theory/polynomial_galois_group.lean | 2 +- src/linear_algebra/matrix/hermitian.lean | 2 +- src/number_theory/number_field/embeddings.lean | 4 ++-- 11 files changed, 24 insertions(+), 24 deletions(-) diff --git a/src/analysis/complex/upper_half_plane/metric.lean b/src/analysis/complex/upper_half_plane/metric.lean index cde0fe244c685..39d5417caffa5 100644 --- a/src/analysis/complex/upper_half_plane/metric.lean +++ b/src/analysis/complex/upper_half_plane/metric.lean @@ -111,7 +111,7 @@ begin div_mul_eq_div_div _ _ (dist _ _), le_div_iff, div_mul_eq_mul_div], { exact div_le_div_of_le (mul_nonneg zero_le_two (sqrt_nonneg _)) (euclidean_geometry.mul_dist_le_mul_dist_add_mul_dist (a : ℂ) b c (conj ↑b)) }, - { rw [dist_comm, dist_pos, ne.def, complex.eq_conj_iff_im], + { rw [dist_comm, dist_pos, ne.def, complex.conj_eq_iff_im], exact b.im_ne_zero } end diff --git a/src/analysis/inner_product_space/positive.lean b/src/analysis/inner_product_space/positive.lean index 3b7e86b17d114..8b35a1d815b43 100644 --- a/src/analysis/inner_product_space/positive.lean +++ b/src/analysis/inner_product_space/positive.lean @@ -119,7 +119,7 @@ lemma is_positive_iff_complex (T : E' →L[ℂ] E') : is_positive T ↔ ∀ x, (re ⟪T x, x⟫_ℂ : ℂ) = ⟪T x, x⟫_ℂ ∧ 0 ≤ re ⟪T x, x⟫_ℂ := begin simp_rw [is_positive, forall_and_distrib, is_self_adjoint_iff_is_symmetric, - linear_map.is_symmetric_iff_inner_map_self_real, eq_conj_iff_re], + linear_map.is_symmetric_iff_inner_map_self_real, conj_eq_iff_re], refl end diff --git a/src/analysis/inner_product_space/spectrum.lean b/src/analysis/inner_product_space/spectrum.lean index c40c794b92592..525c6e21e8b0f 100644 --- a/src/analysis/inner_product_space/spectrum.lean +++ b/src/analysis/inner_product_space/spectrum.lean @@ -205,7 +205,7 @@ begin have H₂ : v ≠ 0 := by simpa using (hT.eigenvector_basis hn).to_basis.ne_zero i, exact ⟨H₁, H₂⟩ }, have re_μ : ↑(is_R_or_C.re μ) = μ, - { rw ← is_R_or_C.eq_conj_iff_re, + { rw ← is_R_or_C.conj_eq_iff_re, exact hT.conj_eigenvalue_eq_self (has_eigenvalue_of_has_eigenvector key) }, simpa [re_μ] using key, end diff --git a/src/analysis/inner_product_space/symmetric.lean b/src/analysis/inner_product_space/symmetric.lean index 51143dd81b7b9..e69fe416d7197 100644 --- a/src/analysis/inner_product_space/symmetric.lean +++ b/src/analysis/inner_product_space/symmetric.lean @@ -109,7 +109,7 @@ end begin rsuffices ⟨r, hr⟩ : ∃ r : ℝ, ⟪T x, x⟫ = r, { simp [hr, T.re_apply_inner_self_apply] }, - rw ← eq_conj_iff_real, + rw ← conj_eq_iff_real, exact hT.conj_inner_sym x x end @@ -165,7 +165,7 @@ begin { simp_rw [h, zero_mul, sub_zero, add_zero, map_add, map_sub, inner_add_left, inner_add_right, inner_sub_left, inner_sub_right, hT x, ← inner_conj_symm x (T y)], suffices : (re ⟪T y, x⟫ : 𝕜) = ⟪T y, x⟫, - { rw eq_conj_iff_re.mpr this, + { rw conj_eq_iff_re.mpr this, ring_nf, }, { rw ← re_add_im ⟪T y, x⟫, simp_rw [h, mul_zero, add_zero], diff --git a/src/analysis/special_functions/gamma.lean b/src/analysis/special_functions/gamma.lean index b47ada18e93b6..fd7163e6326af 100644 --- a/src/analysis/special_functions/gamma.lean +++ b/src/analysis/special_functions/gamma.lean @@ -612,7 +612,7 @@ lemma Gamma_one : Gamma 1 = 1 := by rw [Gamma, complex.of_real_one, complex.Gamma_one, complex.one_re] lemma _root_.complex.Gamma_of_real (s : ℝ) : complex.Gamma (s : ℂ) = Gamma s := -by rw [Gamma, eq_comm, ←complex.eq_conj_iff_re, ←complex.Gamma_conj, complex.conj_of_real] +by rw [Gamma, eq_comm, ←complex.conj_eq_iff_re, ←complex.Gamma_conj, complex.conj_of_real] theorem Gamma_nat_eq_factorial (n : ℕ) : Gamma (n + 1) = n! := by rw [Gamma, complex.of_real_add, complex.of_real_nat_cast, complex.of_real_one, diff --git a/src/data/complex/basic.lean b/src/data/complex/basic.lean index 887274d0b7c8e..74951e27e6bbf 100644 --- a/src/data/complex/basic.lean +++ b/src/data/complex/basic.lean @@ -257,14 +257,14 @@ lemma conj_bit1 (z : ℂ) : conj (bit1 z) = bit1 (conj z) := ext_iff.2 $ by simp @[simp] lemma conj_neg_I : conj (-I) = I := ext_iff.2 $ by simp -lemma eq_conj_iff_real {z : ℂ} : conj z = z ↔ ∃ r : ℝ, z = r := +lemma conj_eq_iff_real {z : ℂ} : conj z = z ↔ ∃ r : ℝ, z = r := ⟨λ h, ⟨z.re, ext rfl $ eq_zero_of_neg_eq (congr_arg im h)⟩, λ ⟨h, e⟩, by rw [e, conj_of_real]⟩ -lemma eq_conj_iff_re {z : ℂ} : conj z = z ↔ (z.re : ℂ) = z := -eq_conj_iff_real.trans ⟨by rintro ⟨r, rfl⟩; simp, λ h, ⟨_, h.symm⟩⟩ +lemma conj_eq_iff_re {z : ℂ} : conj z = z ↔ (z.re : ℂ) = z := +conj_eq_iff_real.trans ⟨by rintro ⟨r, rfl⟩; simp, λ h, ⟨_, h.symm⟩⟩ -lemma eq_conj_iff_im {z : ℂ} : conj z = z ↔ z.im = 0 := +lemma conj_eq_iff_im {z : ℂ} : conj z = z ↔ z.im = 0 := ⟨λ h, add_self_eq_zero.mp (neg_eq_iff_add_eq_zero.mp (congr_arg im h)), λ h, ext rfl (neg_eq_iff_add_eq_zero.mpr (add_self_eq_zero.mpr h))⟩ diff --git a/src/data/complex/exponential.lean b/src/data/complex/exponential.lean index f7a9a272be1af..9591e4896fee2 100644 --- a/src/data/complex/exponential.lean +++ b/src/data/complex/exponential.lean @@ -474,7 +474,7 @@ begin end @[simp] lemma of_real_exp_of_real_re (x : ℝ) : ((exp x).re : ℂ) = exp x := -eq_conj_iff_re.1 $ by rw [← exp_conj, conj_of_real] +conj_eq_iff_re.1 $ by rw [← exp_conj, conj_of_real] @[simp, norm_cast] lemma of_real_exp (x : ℝ) : (real.exp x : ℂ) = exp x := of_real_exp_of_real_re _ @@ -537,7 +537,7 @@ by rw [sinh, ← ring_hom.map_neg, exp_conj, exp_conj, ← ring_hom.map_sub, sin map_div₀, conj_bit0, ring_hom.map_one] @[simp] lemma of_real_sinh_of_real_re (x : ℝ) : ((sinh x).re : ℂ) = sinh x := -eq_conj_iff_re.1 $ by rw [← sinh_conj, conj_of_real] +conj_eq_iff_re.1 $ by rw [← sinh_conj, conj_of_real] @[simp, norm_cast] lemma of_real_sinh (x : ℝ) : (real.sinh x : ℂ) = sinh x := of_real_sinh_of_real_re _ @@ -554,7 +554,7 @@ begin end lemma of_real_cosh_of_real_re (x : ℝ) : ((cosh x).re : ℂ) = cosh x := -eq_conj_iff_re.1 $ by rw [← cosh_conj, conj_of_real] +conj_eq_iff_re.1 $ by rw [← cosh_conj, conj_of_real] @[simp, norm_cast] lemma of_real_cosh (x : ℝ) : (real.cosh x : ℂ) = cosh x := of_real_cosh_of_real_re _ @@ -574,7 +574,7 @@ lemma tanh_conj : tanh (conj x) = conj (tanh x) := by rw [tanh, sinh_conj, cosh_conj, ← map_div₀, tanh] @[simp] lemma of_real_tanh_of_real_re (x : ℝ) : ((tanh x).re : ℂ) = tanh x := -eq_conj_iff_re.1 $ by rw [← tanh_conj, conj_of_real] +conj_eq_iff_re.1 $ by rw [← tanh_conj, conj_of_real] @[simp, norm_cast] lemma of_real_tanh (x : ℝ) : (real.tanh x : ℂ) = tanh x := of_real_tanh_of_real_re _ @@ -757,7 +757,7 @@ by rw [← mul_left_inj' I_ne_zero, ← sinh_mul_I, mul_neg, sinh_neg, sinh_mul_I, mul_neg] @[simp] lemma of_real_sin_of_real_re (x : ℝ) : ((sin x).re : ℂ) = sin x := -eq_conj_iff_re.1 $ by rw [← sin_conj, conj_of_real] +conj_eq_iff_re.1 $ by rw [← sin_conj, conj_of_real] @[simp, norm_cast] lemma of_real_sin (x : ℝ) : (real.sin x : ℂ) = sin x := of_real_sin_of_real_re _ @@ -772,7 +772,7 @@ by rw [← cosh_mul_I, ← conj_neg_I, ← ring_hom.map_mul, ← cosh_mul_I, cosh_conj, mul_neg, cosh_neg] @[simp] lemma of_real_cos_of_real_re (x : ℝ) : ((cos x).re : ℂ) = cos x := -eq_conj_iff_re.1 $ by rw [← cos_conj, conj_of_real] +conj_eq_iff_re.1 $ by rw [← cos_conj, conj_of_real] @[simp, norm_cast] lemma of_real_cos (x : ℝ) : (real.cos x : ℂ) = cos x := of_real_cos_of_real_re _ @@ -795,7 +795,7 @@ lemma tan_conj : tan (conj x) = conj (tan x) := by rw [tan, sin_conj, cos_conj, ← map_div₀, tan] @[simp] lemma of_real_tan_of_real_re (x : ℝ) : ((tan x).re : ℂ) = tan x := -eq_conj_iff_re.1 $ by rw [← tan_conj, conj_of_real] +conj_eq_iff_re.1 $ by rw [← tan_conj, conj_of_real] @[simp, norm_cast] lemma of_real_tan (x : ℝ) : (real.tan x : ℂ) = tan x := of_real_tan_of_real_re _ diff --git a/src/data/is_R_or_C/basic.lean b/src/data/is_R_or_C/basic.lean index d0508ec16e351..dd036822025a4 100644 --- a/src/data/is_R_or_C/basic.lean +++ b/src/data/is_R_or_C/basic.lean @@ -224,7 +224,7 @@ begin simp only [one_mul, algebra.smul_mul_assoc], end -lemma eq_conj_iff_real {z : K} : conj z = z ↔ ∃ r : ℝ, z = (r : K) := +lemma conj_eq_iff_real {z : K} : conj z = z ↔ ∃ r : ℝ, z = (r : K) := begin split, { intro h, @@ -249,8 +249,8 @@ abbreviation conj_to_ring_equiv : K ≃+* Kᵐᵒᵖ := star_ring_equiv variables {K} -lemma eq_conj_iff_re {z : K} : conj z = z ↔ ((re z) : K) = z := -eq_conj_iff_real.trans ⟨by rintro ⟨r, rfl⟩; simp, λ h, ⟨_, h.symm⟩⟩ +lemma conj_eq_iff_re {z : K} : conj z = z ↔ ((re z) : K) = z := +conj_eq_iff_real.trans ⟨by rintro ⟨r, rfl⟩; simp, λ h, ⟨_, h.symm⟩⟩ /-- The norm squared function. -/ def norm_sq : K →*₀ ℝ := diff --git a/src/field_theory/polynomial_galois_group.lean b/src/field_theory/polynomial_galois_group.lean index c6806e3124acf..4c73c3d5bf348 100644 --- a/src/field_theory/polynomial_galois_group.lean +++ b/src/field_theory/polynomial_galois_group.lean @@ -374,7 +374,7 @@ begin (restrict p ℂ (complex.conj_ae.restrict_scalars ℚ)) w = w ↔ w.val.im = 0, { intro w, rw [subtype.ext_iff, gal_action_hom_restrict], - exact complex.eq_conj_iff_im }, + exact complex.conj_eq_iff_im }, have hc : ∀ z : ℂ, z ∈ c ↔ aeval z p = 0 ∧ z.im ≠ 0, { intro z, simp_rw [finset.mem_image, exists_prop], diff --git a/src/linear_algebra/matrix/hermitian.lean b/src/linear_algebra/matrix/hermitian.lean index 9bb0f8bdc780f..b4a8e6b46e124 100644 --- a/src/linear_algebra/matrix/hermitian.lean +++ b/src/linear_algebra/matrix/hermitian.lean @@ -227,7 +227,7 @@ variables [is_R_or_C α] [is_R_or_C β] /-- The diagonal elements of a complex hermitian matrix are real. -/ lemma is_hermitian.coe_re_apply_self {A : matrix n n α} (h : A.is_hermitian) (i : n) : (re (A i i) : α) = A i i := -by rw [←eq_conj_iff_re, ←star_def, ←conj_transpose_apply, h.eq] +by rw [←conj_eq_iff_re, ←star_def, ←conj_transpose_apply, h.eq] /-- The diagonal elements of a complex hermitian matrix are real. -/ lemma is_hermitian.coe_re_diag {A : matrix n n α} (h : A.is_hermitian) : diff --git a/src/number_theory/number_field/embeddings.lean b/src/number_theory/number_field/embeddings.lean index ed90ed59b4282..2e1883a6753cd 100644 --- a/src/number_theory/number_field/embeddings.lean +++ b/src/number_theory/number_field/embeddings.lean @@ -172,7 +172,7 @@ lemma is_real_iff {φ : K →+* ℂ} : is_real φ ↔ conjugate φ = φ := is_se def is_real.embedding {φ : K →+* ℂ} (hφ : is_real φ) : K →+* ℝ := { to_fun := λ x, (φ x).re, map_one' := by simp only [map_one, one_re], - map_mul' := by simp only [complex.eq_conj_iff_im.mp (ring_hom.congr_fun hφ _), map_mul, mul_re, + map_mul' := by simp only [complex.conj_eq_iff_im.mp (ring_hom.congr_fun hφ _), map_mul, mul_re, mul_zero, tsub_zero, eq_self_iff_true, forall_const], map_zero' := by simp only [map_zero, zero_re], map_add' := by simp only [map_add, add_re, eq_self_iff_true, forall_const], } @@ -182,7 +182,7 @@ lemma is_real.coe_embedding_apply {φ : K →+* ℂ} (hφ : is_real φ) (x : K) (hφ.embedding x : ℂ) = φ x := begin ext, { refl, }, - { rw [of_real_im, eq_comm, ← complex.eq_conj_iff_im], + { rw [of_real_im, eq_comm, ← complex.conj_eq_iff_im], rw is_real at hφ, exact ring_hom.congr_fun hφ x, }, end From f9dd3204df14a0749cd456fac1e6849dfe7d2b88 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Wed, 3 May 2023 06:27:57 +0000 Subject: [PATCH 0937/1291] chore(analysis/normed_space/basic): rename `abs_norm_eq_norm` to `abs_norm` (#18921) --- src/analysis/complex/open_mapping.lean | 2 +- src/analysis/inner_product_space/basic.lean | 6 +++--- src/analysis/inner_product_space/calculus.lean | 2 +- src/analysis/normed_space/basic.lean | 8 ++++---- src/analysis/normed_space/lp_space.lean | 2 +- .../normed_space/star/continuous_functional_calculus.lean | 2 +- 6 files changed, 11 insertions(+), 11 deletions(-) diff --git a/src/analysis/complex/open_mapping.lean b/src/analysis/complex/open_mapping.lean index bb944e46be3a0..9e347d273ff1a 100644 --- a/src/analysis/complex/open_mapping.lean +++ b/src/analysis/complex/open_mapping.lean @@ -140,7 +140,7 @@ begin simpa [gray, ray] using h w e2 }, have h4 : ‖z - z₀‖ < r := by simpa [dist_eq_norm] using mem_ball.mp hz, replace h4 : ↑‖z - z₀‖ ∈ ball (0 : ℂ) r := by simpa only [mem_ball_zero_iff, norm_eq_abs, - abs_of_real, abs_norm_eq_norm], + abs_of_real, abs_norm], simpa only [gray, ray, smul_smul, mul_inv_cancel h', one_smul, add_sub_cancel'_right, function.comp_app, coe_smul] using h3 ↑‖z - z₀‖ h4 }, { right, -- Otherwise, it is open along at least one direction and that implies the result diff --git a/src/analysis/inner_product_space/basic.lean b/src/analysis/inner_product_space/basic.lean index 242bb6c7a5b9d..0d1f1aba17dd4 100644 --- a/src/analysis/inner_product_space/basic.lean +++ b/src/analysis/inner_product_space/basic.lean @@ -1519,7 +1519,7 @@ begin is_R_or_C.abs_div, is_R_or_C.abs_mul, ←inner_self_re_abs, inner_self_eq_norm_mul_norm] at h, norm_cast at h, - rwa [_root_.abs_mul, abs_norm_eq_norm, abs_norm_eq_norm, ←mul_assoc, mul_comm, + rwa [_root_.abs_mul, abs_norm, abs_norm, ←mul_assoc, mul_comm, mul_div_mul_left _ _ (λ h, hx0 (norm_eq_zero.1 h)), ←is_R_or_C.norm_eq_abs, ←norm_smul] at h }, have hr0 : r ≠ 0, @@ -1544,7 +1544,7 @@ begin rcases h with ⟨hx, ⟨r, ⟨hr, hy⟩⟩⟩, rw [hy, is_R_or_C.abs_div], norm_cast, - rw [_root_.abs_mul, abs_norm_eq_norm, abs_norm_eq_norm], + rw [_root_.abs_mul, abs_norm, abs_norm], exact abs_inner_div_norm_mul_norm_eq_one_of_ne_zero_of_ne_zero_mul hx hr } end @@ -1648,7 +1648,7 @@ begin ... ↔ ‖(‖y‖:𝕜) • x - (‖x‖:𝕜) • y‖ * ‖(‖y‖:𝕜) • x - (‖x‖:𝕜) • y‖ = 0 : begin simp only [@norm_sub_mul_self 𝕜, inner_smul_left, inner_smul_right, norm_smul, conj_of_real, - is_R_or_C.norm_eq_abs, abs_of_real, of_real_im, of_real_re, mul_re, abs_norm_eq_norm], + is_R_or_C.norm_eq_abs, abs_of_real, of_real_im, of_real_re, mul_re, abs_norm], refine eq.congr _ rfl, ring end diff --git a/src/analysis/inner_product_space/calculus.lean b/src/analysis/inner_product_space/calculus.lean index 4abc6d715a9e2..31ddee695201e 100644 --- a/src/analysis/inner_product_space/calculus.lean +++ b/src/analysis/inner_product_space/calculus.lean @@ -364,7 +364,7 @@ begin refl, }, refine cont_diff_at.congr_of_eventually_eq _ hf, suffices : cont_diff_at ℝ n (λy, (1 - ‖(y : E)‖^2).sqrt⁻¹) y, { exact this.smul cont_diff_at_id }, - have h : 0 < 1 - ‖(y : E)‖^2, by rwa [mem_ball_zero_iff, ← _root_.abs_one, ← abs_norm_eq_norm, + have h : 0 < 1 - ‖(y : E)‖^2, by rwa [mem_ball_zero_iff, ← _root_.abs_one, ← abs_norm, ← sq_lt_sq, one_pow, ← sub_pos] at hy, refine cont_diff_at.inv _ (real.sqrt_ne_zero'.mpr h), refine cont_diff_at.comp _ (cont_diff_at_sqrt h.ne.symm) _, diff --git a/src/analysis/normed_space/basic.lean b/src/analysis/normed_space/basic.lean index 830b4367d1f56..3e70aa5b481c3 100644 --- a/src/analysis/normed_space/basic.lean +++ b/src/analysis/normed_space/basic.lean @@ -94,8 +94,8 @@ lemma norm_zsmul (α) [normed_field α] [normed_space α β] (n : ℤ) (x : β) ‖n • x‖ = ‖(n : α)‖ * ‖x‖ := by rw [← norm_smul, ← int.smul_one_eq_coe, smul_assoc, one_smul] -@[simp] lemma abs_norm_eq_norm (z : β) : |‖z‖| = ‖z‖ := - (abs_eq (norm_nonneg z)).mpr (or.inl rfl) +@[simp] lemma abs_norm (z : β) : |‖z‖| = ‖z‖ := +abs_of_nonneg $ norm_nonneg z lemma inv_norm_smul_mem_closed_unit_ball [normed_space ℝ β] (x : β) : ‖x‖⁻¹ • x ∈ closed_ball (0 : β) 1 := @@ -225,8 +225,8 @@ noncomputable def homeomorph_unit_ball [normed_space ℝ E] : { to_fun := λ x, ⟨(1 + ‖x‖^2).sqrt⁻¹ • x, begin have : 0 < 1 + ‖x‖ ^ 2, by positivity, rw [mem_ball_zero_iff, norm_smul, real.norm_eq_abs, abs_inv, ← div_eq_inv_mul, - div_lt_one (abs_pos.mpr $ real.sqrt_ne_zero'.mpr this), ← abs_norm_eq_norm x, ← sq_lt_sq, - abs_norm_eq_norm, real.sq_sqrt this.le], + div_lt_one (abs_pos.mpr $ real.sqrt_ne_zero'.mpr this), ← abs_norm x, ← sq_lt_sq, + abs_norm, real.sq_sqrt this.le], exact lt_one_add _, end⟩, inv_fun := λ y, (1 - ‖(y : E)‖^2).sqrt⁻¹ • (y : E), diff --git a/src/analysis/normed_space/lp_space.lean b/src/analysis/normed_space/lp_space.lean index 77c3b45e3042f..9cfaeab54c8aa 100644 --- a/src/analysis/normed_space/lp_space.lean +++ b/src/analysis/normed_space/lp_space.lean @@ -961,7 +961,7 @@ begin rw ← H at hs, have : |‖∑ i in s, lp.single p i (f i : E i) - f‖ ^ p.to_real| = ‖∑ i in s, lp.single p i (f i : E i) - f‖ ^ p.to_real, - { simp only [real.abs_rpow_of_nonneg (norm_nonneg _), abs_norm_eq_norm] }, + { simp only [real.abs_rpow_of_nonneg (norm_nonneg _), abs_norm] }, linarith end diff --git a/src/analysis/normed_space/star/continuous_functional_calculus.lean b/src/analysis/normed_space/star/continuous_functional_calculus.lean index adaba66e286e7..df35d3e3b9eb6 100644 --- a/src/analysis/normed_space/star/continuous_functional_calculus.lean +++ b/src/analysis/normed_space/star/continuous_functional_calculus.lean @@ -146,7 +146,7 @@ begin refine ⟨u.unit_of_nearby _ _, rfl⟩, simp only [complex.abs_of_real, map_inv₀, units.coe_map, units.coe_inv, ring_hom.coe_monoid_hom, ring_hom.to_monoid_hom_eq_coe, units.coe_mk0, units.coe_map_inv, norm_algebra_map', - inv_inv, complex.norm_eq_abs, abs_norm_eq_norm, subtype.val_eq_coe, coe_coe], + inv_inv, complex.norm_eq_abs, abs_norm, subtype.val_eq_coe, coe_coe], /- Since `a` is invertible, by `spectrum_star_mul_self_of_is_star_normal`, the spectrum (in `A`) of `star a * a` is contained in the half-open interval `(0, ‖star a * a‖]`. Therefore, by basic spectral mapping properties, the spectrum of `‖star a * a‖ • 1 - star a * a` is contained in From 4d06b17aea8cf2e220f0b0aa46cd0231593c5c97 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Wed, 3 May 2023 06:27:58 +0000 Subject: [PATCH 0938/1291] chore(data/set/intervals/monotone): fix names (#18924) Some lemmas used `Iic` instead of `Ici` in the names, probably because of a copy+paste error. --- src/data/set/intervals/monotone.lean | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/src/data/set/intervals/monotone.lean b/src/data/set/intervals/monotone.lean index 90f52eb82e83a..ee16fb679330a 100644 --- a/src/data/set/intervals/monotone.lean +++ b/src/data/set/intervals/monotone.lean @@ -182,7 +182,7 @@ begin { exact ih (strict_mono_on.mono hφ (λ x hx, le_trans hx (le_succ _))) _ h } end -lemma strict_mono_on.Iic_le_id [pred_order α] [is_pred_archimedean α] [order_top α] +lemma strict_mono_on.Ici_le_id [pred_order α] [is_pred_archimedean α] [order_top α] {n : α} {φ : α → α} (hφ : strict_mono_on φ (set.Ici n)) : ∀ m, n ≤ m → φ m ≤ m := @strict_mono_on.Iic_id_le αᵒᵈ _ _ _ _ _ _ (λ i hi j hj hij, hφ hj hi hij) @@ -221,12 +221,12 @@ lemma strict_anti_on_Iic_of_succ_lt [succ_order α] [is_succ_archimedean α] strict_anti_on ψ (set.Iic n) := λ i hi j hj hij, @strict_mono_on_Iic_of_lt_succ α βᵒᵈ _ _ ψ _ _ n hψ i hi j hj hij -lemma strict_mono_on_Iic_of_pred_lt [pred_order α] [is_pred_archimedean α] +lemma strict_mono_on_Ici_of_pred_lt [pred_order α] [is_pred_archimedean α] {n : α} (hψ : ∀ m, n < m → ψ (pred m) < ψ m) : strict_mono_on ψ (set.Ici n) := λ i hi j hj hij, @strict_mono_on_Iic_of_lt_succ αᵒᵈ βᵒᵈ _ _ ψ _ _ n hψ j hj i hi hij -lemma strict_anti_on_Iic_of_lt_pred [pred_order α] [is_pred_archimedean α] +lemma strict_anti_on_Ici_of_lt_pred [pred_order α] [is_pred_archimedean α] {n : α} (hψ : ∀ m, n < m → ψ m < ψ (pred m)) : strict_anti_on ψ (set.Ici n) := λ i hi j hj hij, @strict_anti_on_Iic_of_succ_lt αᵒᵈ βᵒᵈ _ _ ψ _ _ n hψ j hj i hi hij From ef997baa41b5c428be3fb50089a7139bf4ee886b Mon Sep 17 00:00:00 2001 From: Amelia Livingston <101damnations@github.com> Date: Wed, 3 May 2023 13:04:49 +0000 Subject: [PATCH 0939/1291] feat(representation_theory/group_cohomology/basic): add standard definition of group cohomology (#18341) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We define the complex of inhomogeneous cochains, define group cohomology to be its cohomology, and prove this is isomorphic to the appropriate Ext groups. Co-authored-by: Joël Riou Co-authored-by: Amelia Livingston --- src/algebra/big_operators/fin.lean | 34 ++++ src/data/fin/tuple/basic.lean | 40 ++++ .../group_cohomology/basic.lean | 184 ++++++++++++++++++ .../resolution.lean} | 26 ++- 4 files changed, 279 insertions(+), 5 deletions(-) create mode 100644 src/representation_theory/group_cohomology/basic.lean rename src/representation_theory/{group_cohomology_resolution.lean => group_cohomology/resolution.lean} (96%) diff --git a/src/algebra/big_operators/fin.lean b/src/algebra/big_operators/fin.lean index d0182e6396d4e..9b07547122eb8 100644 --- a/src/algebra/big_operators/fin.lean +++ b/src/algebra/big_operators/fin.lean @@ -208,6 +208,40 @@ begin assoc_rw [hi, inv_mul_cancel_left] } end +/-- Let `(g₀, g₁, ..., gₙ)` be a tuple of elements in `Gⁿ⁺¹`. +Then if `k < j`, this says `(g₀g₁...gₖ₋₁)⁻¹ * g₀g₁...gₖ = gₖ`. +If `k = j`, it says `(g₀g₁...gₖ₋₁)⁻¹ * g₀g₁...gₖ₊₁ = gₖgₖ₊₁`. +If `k > j`, it says `(g₀g₁...gₖ)⁻¹ * g₀g₁...gₖ₊₁ = gₖ₊₁.` +Useful for defining group cohomology. -/ +@[to_additive "Let `(g₀, g₁, ..., gₙ)` be a tuple of elements in `Gⁿ⁺¹`. +Then if `k < j`, this says `-(g₀ + g₁ + ... + gₖ₋₁) + (g₀ + g₁ + ... + gₖ) = gₖ`. +If `k = j`, it says `-(g₀ + g₁ + ... + gₖ₋₁) + (g₀ + g₁ + ... + gₖ₊₁) = gₖ + gₖ₊₁`. +If `k > j`, it says `-(g₀ + g₁ + ... + gₖ) + (g₀ + g₁ + ... + gₖ₊₁) = gₖ₊₁.` +Useful for defining group cohomology."] +lemma inv_partial_prod_mul_eq_contract_nth {G : Type*} [group G] + (g : fin (n + 1) → G) (j : fin (n + 1)) (k : fin n) : + (partial_prod g (j.succ.succ_above k.cast_succ))⁻¹ * partial_prod g (j.succ_above k).succ + = j.contract_nth has_mul.mul g k := +begin + have := partial_prod_right_inv (1 : G) g, + simp only [one_smul, coe_eq_cast_succ] at this, + rcases lt_trichotomy (k : ℕ) j with (h|h|h), + { rwa [succ_above_below, succ_above_below, this, contract_nth_apply_of_lt], + { assumption }, + { rw [cast_succ_lt_iff_succ_le, succ_le_succ_iff, le_iff_coe_le_coe], + exact le_of_lt h }}, + { rwa [succ_above_below, succ_above_above, partial_prod_succ, cast_succ_fin_succ, ←mul_assoc, + this, contract_nth_apply_of_eq], + { simpa only [le_iff_coe_le_coe, ←h] }, + { rw [cast_succ_lt_iff_succ_le, succ_le_succ_iff, le_iff_coe_le_coe], + exact le_of_eq h }}, + { rwa [succ_above_above, succ_above_above, partial_prod_succ, partial_prod_succ, + cast_succ_fin_succ, partial_prod_succ, inv_mul_cancel_left, contract_nth_apply_of_gt], + { exact le_iff_coe_le_coe.2 (le_of_lt h) }, + { rw [le_iff_coe_le_coe, coe_succ], + exact nat.succ_le_of_lt h }}, +end + end partial_prod end fin diff --git a/src/data/fin/tuple/basic.lean b/src/data/fin/tuple/basic.lean index eddc29b8946b3..e9509b715fdca 100644 --- a/src/data/fin/tuple/basic.lean +++ b/src/data/fin/tuple/basic.lean @@ -847,6 +847,46 @@ lemma mem_find_of_unique {p : fin n → Prop} [decidable_pred p] mem_find_iff.2 ⟨hi, λ j hj, le_of_eq $ h i j hi hj⟩ end find +section contract_nth + +variables {α : Type*} + +/-- Sends `(g₀, ..., gₙ)` to `(g₀, ..., op gⱼ gⱼ₊₁, ..., gₙ)`. -/ +def contract_nth (j : fin (n + 1)) (op : α → α → α) (g : fin (n + 1) → α) (k : fin n) : α := +if (k : ℕ) < j then g k.cast_succ else +if (k : ℕ) = j then op (g k.cast_succ) (g k.succ) +else g k.succ + +lemma contract_nth_apply_of_lt (j : fin (n + 1)) (op : α → α → α) (g : fin (n + 1) → α) + (k : fin n) (h : (k : ℕ) < j) : + contract_nth j op g k = g k.cast_succ := if_pos h + +lemma contract_nth_apply_of_eq (j : fin (n + 1)) (op : α → α → α) (g : fin (n + 1) → α) + (k : fin n) (h : (k : ℕ) = j) : + contract_nth j op g k = op (g k.cast_succ) (g k.succ) := +begin + have : ¬(k : ℕ) < j, from not_lt.2 (le_of_eq h.symm), + rw [contract_nth, if_neg this, if_pos h], +end + +lemma contract_nth_apply_of_gt (j : fin (n + 1)) (op : α → α → α) (g : fin (n + 1) → α) + (k : fin n) (h : (j : ℕ) < k) : + contract_nth j op g k = g k.succ := +by rw [contract_nth, if_neg (not_lt_of_gt h), if_neg (ne.symm $ ne_of_lt h)] + +lemma contract_nth_apply_of_ne (j : fin (n + 1)) (op : α → α → α) (g : fin (n + 1) → α) + (k : fin n) (hjk : (j : ℕ) ≠ k) : + contract_nth j op g k = g (j.succ_above k) := +begin + rcases lt_trichotomy (k : ℕ) j with (h|h|h), + { rwa [j.succ_above_below, contract_nth_apply_of_lt], + { rwa [ fin.lt_iff_coe_lt_coe] }}, + { exact false.elim (hjk h.symm) }, + { rwa [j.succ_above_above, contract_nth_apply_of_gt], + { exact fin.le_iff_coe_le_coe.2 (le_of_lt h) }} +end + +end contract_nth /-- To show two sigma pairs of tuples agree, it to show the second elements are related via `fin.cast`. -/ diff --git a/src/representation_theory/group_cohomology/basic.lean b/src/representation_theory/group_cohomology/basic.lean new file mode 100644 index 0000000000000..e63ce48b3af5c --- /dev/null +++ b/src/representation_theory/group_cohomology/basic.lean @@ -0,0 +1,184 @@ +/- +Copyright (c) 2023 Amelia Livingston. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Amelia Livingston +-/ + +import algebra.homology.opposite +import representation_theory.group_cohomology.resolution + +/-! +# The group cohomology of a `k`-linear `G`-representation + +Let `k` be a commutative ring and `G` a group. This file defines the group cohomology of +`A : Rep k G` to be the cohomology of the complex +$$0 \to \mathrm{Fun}(G^0, A) \to \mathrm{Fun}(G^1, A) \to \mathrm{Fun}(G^2, A) \to \dots$$ +with differential $d^n$ sending $f: G^n \to A$ to the function mapping $(g_0, \dots, g_n)$ to +$$\rho(g_0)(f(g_1, \dots, g_n)) ++ \sum_{i = 0}^{n - 1} (-1)^{i + 1}\cdot f(g_0, \dots, g_ig_{i + 1}, \dots, g_n)$$ +$$+ (-1)^{n + 1}\cdot f(g_0, \dots, g_{n - 1})$$ (where `ρ` is the representation attached to `A`). + +We have a `k`-linear isomorphism $\mathrm{Fun}(G^n, A) \cong \mathrm{Hom}(k[G^{n + 1}], A)$, where +the righthand side is morphisms in `Rep k G`, and the representation on $k[G^{n + 1}]$ +is induced by the diagonal action of `G`. If we conjugate the $n$th differential in +$\mathrm{Hom}(P, A)$ by this isomorphism, where `P` is the standard resolution of `k` as a trivial +`k`-linear `G`-representation, then the resulting map agrees with the differential $d^n$ defined +above, a fact we prove. + +This gives us for free a proof that our $d^n$ squares to zero. It also gives us an isomorphism +$\mathrm{H}^n(G, A) \cong \mathrm{Ext}^n(k, A),$ where $\mathrm{Ext}$ is taken in the category +`Rep k G`. + +## Main definitions + +* `group_cohomology.linear_yoneda_obj_resolution A`: a complex whose objects are the representation +morphisms $\mathrm{Hom}(k[G^{n + 1}], A)$ and whose cohomology is the group cohomology +$\mathrm{H}^n(G, A)$. +* `group_cohomology.inhomogeneous_cochains A`: a complex whose objects are +$\mathrm{Fun}(G^n, A)$ and whose cohomology is the group cohomology $\mathrm{H}^n(G, A).$ +* `group_cohomology.inhomogeneous_cochains_iso A`: an isomorphism between the above two complexes. +* `group_cohomology A n`: this is $\mathrm{H}^n(G, A),$ defined as the $n$th cohomology of the +second complex, `inhomogeneous_cochains A`. +* `group_cohomology_iso_Ext A n`: an isomorphism $\mathrm{H}^n(G, A) \cong \mathrm{Ext}^n(k, A)$ +(where $\mathrm{Ext}$ is taken in the category `Rep k G`) induced by `inhomogeneous_cochains_iso A`. + +## Implementation notes + +Group cohomology is typically stated for `G`-modules, or equivalently modules over the group ring +`ℤ[G].` However, `ℤ` can be generalized to any commutative ring `k`, which is what we use. +Moreover, we express `k[G]`-module structures on a module `k`-module `A` using the `Rep` +definition. We avoid using instances `module (monoid_algebra k G) A` so that we do not run into +possible scalar action diamonds. + +## TODO + +* API for cohomology in low degree: $\mathrm{H}^0, \mathrm{H}^1$ and $\mathrm{H}^2.$ For example, +the inflation-restriction exact sequence. +* The long exact sequence in cohomology attached to a short exact sequence of representations. +* Upgrading `group_cohomology_iso_Ext` to an isomorphism of derived functors. +* Profinite cohomology. + +Longer term: +* The Hochschild-Serre spectral sequence (this is perhaps a good toy example for the theory of +spectral sequences in general). +-/ + +noncomputable theory +universes u + +variables {k G : Type u} [comm_ring k] {n : ℕ} + +open category_theory +namespace group_cohomology +variables [monoid G] + +/-- The complex `Hom(P, A)`, where `P` is the standard resolution of `k` as a trivial `k`-linear +`G`-representation. -/ +abbreviation linear_yoneda_obj_resolution (A : Rep k G) : cochain_complex (Module.{u} k) ℕ := +homological_complex.unop + ((((linear_yoneda k (Rep k G)).obj A).right_op.map_homological_complex _).obj (resolution k G)) + +lemma linear_yoneda_obj_resolution_d_apply {A : Rep k G} (i j : ℕ) (x : (resolution k G).X i ⟶ A) : + (linear_yoneda_obj_resolution A).d i j x = (resolution k G).d j i ≫ x := +rfl + +end group_cohomology +namespace inhomogeneous_cochains +open Rep group_cohomology + +/-- The differential in the complex of inhomogeneous cochains used to +calculate group cohomology. -/ +@[simps] def d [monoid G] (n : ℕ) (A : Rep k G) : + ((fin n → G) → A) →ₗ[k] (fin (n + 1) → G) → A := +{ to_fun := λ f g, A.ρ (g 0) (f (λ i, g i.succ)) + + finset.univ.sum (λ j : fin (n + 1), (-1 : k) ^ ((j : ℕ) + 1) • f (fin.contract_nth j (*) g)), + map_add' := λ f g, + begin + ext x, + simp only [pi.add_apply, map_add, smul_add, finset.sum_add_distrib, add_add_add_comm], + end, + map_smul' := λ r f, + begin + ext x, + simp only [pi.smul_apply, ring_hom.id_apply, map_smul, smul_add, finset.smul_sum, + ←smul_assoc, smul_eq_mul, mul_comm r], + end } + +variables [group G] (n) (A : Rep k G) + +/-- The theorem that our isomorphism `Fun(Gⁿ, A) ≅ Hom(k[Gⁿ⁺¹], A)` (where the righthand side is +morphisms in `Rep k G`) commutes with the differentials in the complex of inhomogeneous cochains +and the homogeneous `linear_yoneda_obj_resolution`. -/ +lemma d_eq : + d n A = ((diagonal_hom_equiv n A).to_Module_iso.inv + ≫ (linear_yoneda_obj_resolution A).d n (n + 1) + ≫ (diagonal_hom_equiv (n + 1) A).to_Module_iso.hom) := +begin + ext f g, + simp only [Module.coe_comp, linear_equiv.coe_coe, function.comp_app, + linear_equiv.to_Module_iso_inv, linear_yoneda_obj_resolution_d_apply, + linear_equiv.to_Module_iso_hom, diagonal_hom_equiv_apply, Action.comp_hom, + resolution.d_eq k G n, resolution.d_of (fin.partial_prod g), linear_map.map_sum, + ←finsupp.smul_single_one _ ((-1 : k) ^ _), map_smul, d_apply], + simp only [@fin.sum_univ_succ _ _ (n + 1), fin.coe_zero, pow_zero, one_smul, fin.succ_above_zero, + diagonal_hom_equiv_symm_apply f (fin.partial_prod g ∘ @fin.succ (n + 1)), function.comp_app, + fin.partial_prod_succ, fin.cast_succ_zero, fin.partial_prod_zero, one_mul], + congr' 1, + { congr, + ext, + have := fin.partial_prod_right_inv (1 : G) g (fin.cast_succ x), + simp only [mul_inv_rev, fin.coe_eq_cast_succ, one_smul, fin.cast_succ_fin_succ] at *, + rw [mul_assoc, ←mul_assoc _ _ (g x.succ), this, inv_mul_cancel_left] }, + { exact finset.sum_congr rfl (λ j hj, + by rw [diagonal_hom_equiv_symm_partial_prod_succ, fin.coe_succ]) } +end + +end inhomogeneous_cochains +namespace group_cohomology +variables [group G] (n) (A : Rep k G) + +open inhomogeneous_cochains + +/-- Given a `k`-linear `G`-representation `A`, this is the complex of inhomogeneous cochains +$$0 \to \mathrm{Fun}(G^0, A) \to \mathrm{Fun}(G^1, A) \to \mathrm{Fun}(G^2, A) \to \dots$$ +which calculates the group cohomology of `A`. -/ +noncomputable abbreviation inhomogeneous_cochains : cochain_complex (Module k) ℕ := +cochain_complex.of (λ n, Module.of k ((fin n → G) → A)) +(λ n, inhomogeneous_cochains.d n A) (λ n, +begin + ext x y, + have := linear_map.ext_iff.1 ((linear_yoneda_obj_resolution A).d_comp_d n (n + 1) (n + 2)), + simp only [Module.coe_comp, function.comp_app] at this, + simp only [Module.coe_comp, function.comp_app, d_eq, linear_equiv.to_Module_iso_hom, + linear_equiv.to_Module_iso_inv, linear_equiv.coe_coe, linear_equiv.symm_apply_apply, + this, linear_map.zero_apply, map_zero, pi.zero_apply], +end) + +/-- Given a `k`-linear `G`-representation `A`, the complex of inhomogeneous cochains is isomorphic +to `Hom(P, A)`, where `P` is the standard resolution of `k` as a trivial `G`-representation. -/ +def inhomogeneous_cochains_iso : + inhomogeneous_cochains A ≅ linear_yoneda_obj_resolution A := +homological_complex.hom.iso_of_components + (λ i, (Rep.diagonal_hom_equiv i A).to_Module_iso.symm) $ +begin + rintros i j (h : i + 1 = j), + subst h, + simp only [cochain_complex.of_d, d_eq, category.assoc, iso.symm_hom, + iso.hom_inv_id, category.comp_id], +end + +end group_cohomology +open group_cohomology + +/-- The group cohomology of a `k`-linear `G`-representation `A`, as the cohomology of its complex +of inhomogeneous cochains. -/ +def group_cohomology [group G] (A : Rep k G) (n : ℕ) : Module k := +(inhomogeneous_cochains A).homology n + +/-- The `n`th group cohomology of a `k`-linear `G`-representation `A` is isomorphic to +`Extⁿ(k, A)` (taken in `Rep k G`), where `k` is a trivial `k`-linear `G`-representation. -/ +def group_cohomology_iso_Ext [group G] (A : Rep k G) (n : ℕ) : + group_cohomology A n ≅ ((Ext k (Rep k G) n).obj + (opposite.op $ Rep.trivial k G k)).obj A := +(homology_obj_iso_of_homotopy_equiv (homotopy_equiv.of_iso (inhomogeneous_cochains_iso _)) _) + ≪≫ (homological_complex.homology_unop _ _) ≪≫ (Ext_iso k G A n).symm diff --git a/src/representation_theory/group_cohomology_resolution.lean b/src/representation_theory/group_cohomology/resolution.lean similarity index 96% rename from src/representation_theory/group_cohomology_resolution.lean rename to src/representation_theory/group_cohomology/resolution.lean index aae7db290557f..d6049837d2f81 100644 --- a/src/representation_theory/group_cohomology_resolution.lean +++ b/src/representation_theory/group_cohomology/resolution.lean @@ -232,25 +232,25 @@ module.free.of_basis (of_mul_action_basis k G n) end basis end group_cohomology.resolution namespace Rep -variables (n) [group G] +variables (n) [group G] (A : Rep k G) open group_cohomology.resolution /-- Given a `k`-linear `G`-representation `A`, the set of representation morphisms `Hom(k[Gⁿ⁺¹], A)` is `k`-linearly isomorphic to the set of functions `Gⁿ → A`. -/ -noncomputable def diagonal_hom_equiv (A : Rep k G) : +noncomputable def diagonal_hom_equiv : (Rep.of_mul_action k G (fin (n + 1) → G) ⟶ A) ≃ₗ[k] ((fin n → G) → A) := linear.hom_congr k ((diagonal_succ k G n).trans ((representation.of_mul_action k G G).Rep_of_tprod_iso 1)) (iso.refl _) ≪≫ₗ ((Rep.monoidal_closed.linear_hom_equiv_comm _ _ _) ≪≫ₗ (Rep.left_regular_hom_equiv _)) ≪≫ₗ (finsupp.llift A k k (fin n → G)).symm -variables {n} +variables {n A} /-- Given a `k`-linear `G`-representation `A`, `diagonal_hom_equiv` is a `k`-linear isomorphism of the set of representation morphisms `Hom(k[Gⁿ⁺¹], A)` with `Fun(Gⁿ, A)`. This lemma says that this sends a morphism of representations `f : k[Gⁿ⁺¹] ⟶ A` to the function `(g₁, ..., gₙ) ↦ f(1, g₁, g₁g₂, ..., g₁g₂...gₙ).` -/ -lemma diagonal_hom_equiv_apply {A : Rep k G} (f : Rep.of_mul_action k G (fin (n + 1) → G) ⟶ A) +lemma diagonal_hom_equiv_apply (f : Rep.of_mul_action k G (fin (n + 1) → G) ⟶ A) (x : fin n → G) : diagonal_hom_equiv n A f x = f.hom (finsupp.single (fin.partial_prod x) 1) := begin unfold diagonal_hom_equiv, @@ -266,7 +266,7 @@ the set of representation morphisms `Hom(k[Gⁿ⁺¹], A)` with `Fun(Gⁿ, A)`. inverse map sends a function `f : Gⁿ → A` to the representation morphism sending `(g₀, ... gₙ) ↦ ρ(g₀)(f(g₀⁻¹g₁, g₁⁻¹g₂, ..., gₙ₋₁⁻¹gₙ))`, where `ρ` is the representation attached to `A`. -/ -lemma diagonal_hom_equiv_symm_apply {A : Rep k G} (f : (fin n → G) → A) (x : fin (n + 1) → G) : +lemma diagonal_hom_equiv_symm_apply (f : (fin n → G) → A) (x : fin (n + 1) → G) : ((diagonal_hom_equiv n A).symm f).hom (finsupp.single x 1) = A.ρ (x 0) (f (λ (i : fin n), (x ↑i)⁻¹ * x i.succ)) := begin @@ -281,6 +281,22 @@ begin finsupp.llift_apply A k k], end +/-- Auxiliary lemma for defining group cohomology, used to show that the isomorphism +`diagonal_hom_equiv` commutes with the differentials in two complexes which compute +group cohomology. -/ +lemma diagonal_hom_equiv_symm_partial_prod_succ + (f : (fin n → G) → A) (g : fin (n + 1) → G) (a : fin (n + 1)) : + ((diagonal_hom_equiv n A).symm f).hom (finsupp.single (fin.partial_prod g ∘ a.succ.succ_above) 1) + = f (fin.contract_nth a (*) g) := +begin + simp only [diagonal_hom_equiv_symm_apply, function.comp_app, fin.succ_succ_above_zero, + fin.partial_prod_zero, map_one, fin.coe_eq_cast_succ, fin.succ_succ_above_succ, + linear_map.one_apply, fin.partial_prod_succ], + congr, + ext, + rw [←fin.partial_prod_succ, fin.inv_partial_prod_mul_eq_contract_nth], +end + end Rep variables (G) From c163ec99dfc664628ca15d215fce0a5b9c265b68 Mon Sep 17 00:00:00 2001 From: "Filippo A. E. Nuccio" Date: Wed, 3 May 2023 16:34:41 +0000 Subject: [PATCH 0940/1291] chore(ring_theory/discrete_valuation_ring): create folder ring_theory/discrete_valuation_ring (#18867) Moved the file `ring_theory/discrete_valuation_ring.lean` inside the folder and renamed it as `basic.lean`; also, added the file `valuation.tfae` giving equivalent conditions for being a dvr to the folder. --- src/number_theory/padics/padic_integers.lean | 2 +- src/ring_theory/dedekind_domain/dvr.lean | 4 ++-- .../basic.lean} | 0 .../{valuation => discrete_valuation_ring}/tfae.lean | 0 src/ring_theory/valuation/valuation_ring.lean | 2 +- src/ring_theory/witt_vector/discrete_valuation_ring.lean | 2 +- 6 files changed, 5 insertions(+), 5 deletions(-) rename src/ring_theory/{discrete_valuation_ring.lean => discrete_valuation_ring/basic.lean} (100%) rename src/ring_theory/{valuation => discrete_valuation_ring}/tfae.lean (100%) diff --git a/src/number_theory/padics/padic_integers.lean b/src/number_theory/padics/padic_integers.lean index 89711aa75ff29..83af846ac091a 100644 --- a/src/number_theory/padics/padic_integers.lean +++ b/src/number_theory/padics/padic_integers.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Robert Y. Lewis, Mario Carneiro, Johan Commelin -/ import number_theory.padics.padic_numbers -import ring_theory.discrete_valuation_ring +import ring_theory.discrete_valuation_ring.basic import topology.metric_space.cau_seq_filter /-! diff --git a/src/ring_theory/dedekind_domain/dvr.lean b/src/ring_theory/dedekind_domain/dvr.lean index 4d4397ea2e091..64659de77d46e 100644 --- a/src/ring_theory/dedekind_domain/dvr.lean +++ b/src/ring_theory/dedekind_domain/dvr.lean @@ -4,10 +4,10 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenji Nakagawa, Anne Baanen, Filippo A. E. Nuccio -/ import ring_theory.dedekind_domain.ideal -import ring_theory.discrete_valuation_ring +import ring_theory.discrete_valuation_ring.tfae import ring_theory.localization.at_prime import ring_theory.localization.submodule -import ring_theory.valuation.tfae + /-! # Dedekind domains diff --git a/src/ring_theory/discrete_valuation_ring.lean b/src/ring_theory/discrete_valuation_ring/basic.lean similarity index 100% rename from src/ring_theory/discrete_valuation_ring.lean rename to src/ring_theory/discrete_valuation_ring/basic.lean diff --git a/src/ring_theory/valuation/tfae.lean b/src/ring_theory/discrete_valuation_ring/tfae.lean similarity index 100% rename from src/ring_theory/valuation/tfae.lean rename to src/ring_theory/discrete_valuation_ring/tfae.lean diff --git a/src/ring_theory/valuation/valuation_ring.lean b/src/ring_theory/valuation/valuation_ring.lean index 979cc5c29a425..e52900b6e6bfa 100644 --- a/src/ring_theory/valuation/valuation_ring.lean +++ b/src/ring_theory/valuation/valuation_ring.lean @@ -7,7 +7,7 @@ import ring_theory.valuation.integers import ring_theory.ideal.local_ring import ring_theory.localization.fraction_ring import ring_theory.localization.integer -import ring_theory.discrete_valuation_ring +import ring_theory.discrete_valuation_ring.basic import ring_theory.bezout import tactic.field_simp diff --git a/src/ring_theory/witt_vector/discrete_valuation_ring.lean b/src/ring_theory/witt_vector/discrete_valuation_ring.lean index b21fa0fa572e6..abe97d3cecbc2 100644 --- a/src/ring_theory/witt_vector/discrete_valuation_ring.lean +++ b/src/ring_theory/witt_vector/discrete_valuation_ring.lean @@ -6,7 +6,7 @@ Authors: Robert Y. Lewis, Heather Macbeth, Johan Commelin import ring_theory.witt_vector.domain import ring_theory.witt_vector.mul_coeff -import ring_theory.discrete_valuation_ring +import ring_theory.discrete_valuation_ring.basic import tactic.linear_combination /-! From 726d2fe4fd196a6e796d3bc827d025abd32673ff Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Wed, 3 May 2023 18:24:34 +0000 Subject: [PATCH 0941/1291] feat(measure_theory/constructions/prod): marginal measures (#18915) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit For `ρ : measure (α × β)`, define `ρ.fst : measure α := ρ.map prod.fst`, and define `ρ.snd` similarly. --- src/measure_theory/constructions/prod.lean | 39 ++++++++++++++++++++++ 1 file changed, 39 insertions(+) diff --git a/src/measure_theory/constructions/prod.lean b/src/measure_theory/constructions/prod.lean index 43b1c74bd77e6..642c354a80fb6 100644 --- a/src/measure_theory/constructions/prod.lean +++ b/src/measure_theory/constructions/prod.lean @@ -1102,4 +1102,43 @@ lemma set_integral_prod_mul {L : Type*} [is_R_or_C L] ∫ z in s ×ˢ t, f z.1 * g z.2 ∂(μ.prod ν) = (∫ x in s, f x ∂μ) * (∫ y in t, g y ∂ν) := by simp only [← measure.prod_restrict s t, integrable_on, integral_prod_mul] +/-! ### Marginals of a measure defined on a product -/ + +namespace measure + +variables {ρ : measure (α × β)} + +/-- Marginal measure on `α` obtained from a measure `ρ` on `α × β`, defined by `ρ.map prod.fst`. -/ +noncomputable +def fst (ρ : measure (α × β)) : measure α := ρ.map prod.fst + +lemma fst_apply {s : set α} (hs : measurable_set s) : ρ.fst s = ρ (prod.fst ⁻¹' s) := +by rw [fst, measure.map_apply measurable_fst hs] + +lemma fst_univ : ρ.fst univ = ρ univ := +by rw [fst_apply measurable_set.univ, preimage_univ] + +instance [is_finite_measure ρ] : is_finite_measure ρ.fst := by { rw fst, apply_instance, } + +instance [is_probability_measure ρ] : is_probability_measure ρ.fst := +{ measure_univ := by { rw fst_univ, exact measure_univ, } } + +/-- Marginal measure on `β` obtained from a measure on `ρ` `α × β`, defined by `ρ.map prod.snd`. -/ +noncomputable +def snd (ρ : measure (α × β)) : measure β := ρ.map prod.snd + +lemma snd_apply {s : set β} (hs : measurable_set s) : ρ.snd s = ρ (prod.snd ⁻¹' s) := +by rw [snd, measure.map_apply measurable_snd hs] + +lemma snd_univ : ρ.snd univ = ρ univ := +by rw [snd_apply measurable_set.univ, preimage_univ] + +instance [is_finite_measure ρ] : is_finite_measure ρ.snd := by { rw snd, apply_instance, } + +instance [is_probability_measure ρ] : is_probability_measure ρ.snd := +{ measure_univ := by { rw snd_univ, exact measure_univ, } } + +end measure + + end measure_theory From 92c69b77c5a7dc0f7eeddb552508633305157caa Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Wed, 3 May 2023 18:24:36 +0000 Subject: [PATCH 0942/1291] chore(analysis/inner_product_space/basic): golf, add/merge lemmas (#18928) - Merge `inner_product_space.of_core.inner_self_nonneg_im` and `inner_product_space.of_core.inner_self_im_zero` into `inner_product_space.of_core.inner_self_im`. - Rename `inner_product_space.of_core.inner_abs_conj_symm` to `inner_product_space.of_core.abs_inner_symm`. - Rename `inner_abs_conj_symm` to `abs_inner_symm`. - Add `inner_product_space.of_core.norm_sq_eq_zero`. - Merge `inner_self_nonneg_im` and `inner_self_im_zero` into `inner_self_im`. - Reorder, golf. --- src/analysis/inner_product_space/basic.lean | 160 ++++++-------------- 1 file changed, 47 insertions(+), 113 deletions(-) diff --git a/src/analysis/inner_product_space/basic.lean b/src/analysis/inner_product_space/basic.lean index 0d1f1aba17dd4..0c771833488a5 100644 --- a/src/analysis/inner_product_space/basic.lean +++ b/src/analysis/inner_product_space/basic.lean @@ -162,12 +162,9 @@ lemma inner_conj_symm (x y : F) : ⟪y, x⟫† = ⟪x, y⟫ := c.conj_symm x y lemma inner_self_nonneg {x : F} : 0 ≤ re ⟪x, x⟫ := c.nonneg_re _ -lemma inner_self_nonneg_im (x : F) : im ⟪x, x⟫ = 0 := +lemma inner_self_im (x : F) : im ⟪x, x⟫ = 0 := by rw [← @of_real_inj 𝕜, im_eq_conj_sub]; simp [inner_conj_symm] -lemma inner_self_im_zero (x : F) : im ⟪x, x⟫ = 0 := -inner_self_nonneg_im _ - lemma inner_add_left (x y z : F) : ⟪x + y, z⟫ = ⟪x, z⟫ + ⟪y, z⟫ := c.add_left _ _ _ @@ -177,7 +174,7 @@ by rw [←inner_conj_symm, inner_add_left, ring_hom.map_add]; simp only [inner_c lemma inner_norm_sq_eq_inner_self (x : F) : (norm_sqF x : 𝕜) = ⟪x, x⟫ := begin rw ext_iff, - exact ⟨by simp only [of_real_re]; refl, by simp only [inner_self_nonneg_im, of_real_im]⟩ + exact ⟨by simp only [of_real_re]; refl, by simp only [inner_self_im, of_real_im]⟩ end lemma inner_re_symm (x y : F) : re ⟪x, y⟫ = re ⟪y, x⟫ := @@ -199,15 +196,19 @@ lemma inner_zero_right (x : F) : ⟪x, 0⟫ = 0 := by rw [←inner_conj_symm, inner_zero_left]; simp only [ring_hom.map_zero] lemma inner_self_eq_zero {x : F} : ⟪x, x⟫ = 0 ↔ x = 0 := -iff.intro (c.definite _) (by { rintro rfl, exact inner_zero_left _ }) +⟨c.definite _, by { rintro rfl, exact inner_zero_left _ }⟩ + +lemma norm_sq_eq_zero {x : F} : norm_sqF x = 0 ↔ x = 0 := +iff.trans (by simp only [norm_sq, ext_iff, map_zero, inner_self_im, eq_self_iff_true, and_true]) + (@inner_self_eq_zero 𝕜 _ _ _ _ _ x) lemma inner_self_ne_zero {x : F} : ⟪x, x⟫ ≠ 0 ↔ x ≠ 0 := inner_self_eq_zero.not lemma inner_self_re_to_K (x : F) : (re ⟪x, x⟫ : 𝕜) = ⟪x, x⟫ := -by norm_num [ext_iff, inner_self_nonneg_im] +by norm_num [ext_iff, inner_self_im] -lemma inner_abs_conj_symm (x y : F) : abs ⟪x, y⟫ = abs ⟪y, x⟫ := +lemma abs_inner_symm (x y : F) : abs ⟪x, y⟫ = abs ⟪y, x⟫ := by rw [←inner_conj_symm, abs_conj] lemma inner_neg_left (x y : F) : ⟪-x, y⟫ = -⟪x, y⟫ := @@ -259,7 +260,7 @@ begin apply hy', rw ext_iff, exact ⟨by simp only [H, zero_re'], - by simp only [inner_self_nonneg_im, add_monoid_hom.map_zero]⟩ }, + by simp only [inner_self_im, add_monoid_hom.map_zero]⟩ }, have h₆ : re ⟪y, y⟫ ≠ 0 := ne_of_gt h₅, have hmain := calc 0 ≤ re ⟪x - T • y, x - T • y⟫ @@ -311,7 +312,7 @@ begin rw H, conv begin - to_lhs, congr, rw [inner_abs_conj_symm], + to_lhs, congr, rw [abs_inner_symm], end, exact inner_mul_inner_self_le y x, end @@ -332,11 +333,7 @@ add_group_norm.to_normed_add_comm_group linarith }, exact nonneg_le_nonneg_of_sq_le_sq (add_nonneg (sqrt_nonneg _) (sqrt_nonneg _)) this, end, - eq_zero_of_map_eq_zero' := λ x hx, (inner_self_eq_zero : ⟪x, x⟫ = 0 ↔ x = 0).1 $ begin - change sqrt (re ⟪x, x⟫) = 0 at hx, - rw [sqrt_eq_zero inner_self_nonneg] at hx, - exact ext (by simp [hx]) (by simp [inner_self_im_zero]), - end } + eq_zero_of_map_eq_zero' := λ x hx, norm_sq_eq_zero.1 $ (sqrt_eq_zero inner_self_nonneg).1 hx } local attribute [instance] to_normed_add_comm_group @@ -395,11 +392,9 @@ lemma real_inner_comm (x y : F) : ⟪y, x⟫_ℝ = ⟪x, y⟫_ℝ := @inner_conj lemma inner_eq_zero_symm {x y : E} : ⟪x, y⟫ = 0 ↔ ⟪y, x⟫ = 0 := ⟨λ h, by simp [←inner_conj_symm, h], λ h, by simp [←inner_conj_symm, h]⟩ -@[simp] lemma inner_self_nonneg_im (x : E) : im ⟪x, x⟫ = 0 := +@[simp] lemma inner_self_im (x : E) : im ⟪x, x⟫ = 0 := by rw [← @of_real_inj 𝕜, im_eq_conj_sub]; simp -lemma inner_self_im_zero (x : E) : im ⟪x, x⟫ = 0 := inner_self_nonneg_im _ - lemma inner_add_left (x y z : E) : ⟪x + y, z⟫ = ⟪x, z⟫ + ⟪y, z⟫ := inner_product_space.add_left _ _ _ @@ -497,47 +492,11 @@ lemma inner_self_nonneg {x : E} : 0 ≤ re ⟪x, x⟫ := by rw [←norm_sq_eq_inner]; exact pow_nonneg (norm_nonneg x) 2 lemma real_inner_self_nonneg {x : F} : 0 ≤ ⟪x, x⟫_ℝ := @inner_self_nonneg ℝ F _ _ _ x -@[simp] lemma inner_self_eq_zero {x : E} : ⟪x, x⟫ = 0 ↔ x = 0 := -begin - split, - { intro h, - have h₁ : re ⟪x, x⟫ = 0 := - by rw is_R_or_C.ext_iff at h; simp only [h.1, zero_re'], - rw [←norm_sq_eq_inner x] at h₁, - rw [←norm_eq_zero], - exact pow_eq_zero h₁ }, - { rintro rfl, - exact inner_zero_left _ } -end - -lemma inner_self_ne_zero {x : E} : ⟪x, x⟫ ≠ 0 ↔ x ≠ 0 := -inner_self_eq_zero.not - -@[simp] lemma inner_self_nonpos {x : E} : re ⟪x, x⟫ ≤ 0 ↔ x = 0 := -begin - split, - { intro h, - rw ←@inner_self_eq_zero 𝕜, - have H₁ : re ⟪x, x⟫ ≥ 0, exact inner_self_nonneg, - have H₂ : re ⟪x, x⟫ = 0, exact le_antisymm h H₁, - rw is_R_or_C.ext_iff, - exact ⟨by simp [H₂], by simp [inner_self_nonneg_im]⟩ }, - { rintro rfl, - simp only [inner_zero_left, add_monoid_hom.map_zero] } -end - -lemma real_inner_self_nonpos {x : F} : ⟪x, x⟫_ℝ ≤ 0 ↔ x = 0 := -by { have h := @inner_self_nonpos ℝ F _ _ _ x, simpa using h } - @[simp] lemma inner_self_re_to_K (x : E) : (re ⟪x, x⟫ : 𝕜) = ⟪x, x⟫ := -is_R_or_C.ext_iff.2 ⟨by simp only [of_real_re], by simp only [inner_self_nonneg_im, of_real_im]⟩ +is_R_or_C.ext_iff.2 ⟨by simp only [of_real_re], by simp only [inner_self_im, of_real_im]⟩ lemma inner_self_eq_norm_sq_to_K (x : E) : ⟪x, x⟫ = (‖x‖ ^ 2 : 𝕜) := -begin - suffices : (is_R_or_C.re ⟪x, x⟫ : 𝕜) = ‖x‖ ^ 2, - { simpa only [inner_self_re_to_K] using this }, - exact_mod_cast (norm_sq_eq_inner x).symm -end +by rw [← inner_self_re_to_K, ← norm_sq_eq_inner, of_real_pow] lemma inner_self_re_abs (x : E) : re ⟪x, x⟫ = abs ⟪x, x⟫ := begin @@ -549,10 +508,22 @@ end lemma inner_self_abs_to_K (x : E) : (absK ⟪x, x⟫ : 𝕜) = ⟪x, x⟫ := by { rw [←inner_self_re_abs], exact inner_self_re_to_K _ } +@[simp] lemma inner_self_eq_zero {x : E} : ⟪x, x⟫ = 0 ↔ x = 0 := +by rw [inner_self_eq_norm_sq_to_K, sq_eq_zero_iff, of_real_eq_zero, norm_eq_zero] + +lemma inner_self_ne_zero {x : E} : ⟪x, x⟫ ≠ 0 ↔ x ≠ 0 := +inner_self_eq_zero.not + +@[simp] lemma inner_self_nonpos {x : E} : re ⟪x, x⟫ ≤ 0 ↔ x = 0 := +by rw [← norm_sq_eq_inner, (sq_nonneg _).le_iff_eq, sq_eq_zero_iff, norm_eq_zero] + +lemma real_inner_self_nonpos {x : F} : ⟪x, x⟫_ℝ ≤ 0 ↔ x = 0 := +@inner_self_nonpos ℝ F _ _ _ x + lemma real_inner_self_abs (x : F) : absR ⟪x, x⟫_ℝ = ⟪x, x⟫_ℝ := by { have h := @inner_self_abs_to_K ℝ F _ _ _ x, simpa using h } -lemma inner_abs_conj_symm (x y : E) : abs ⟪x, y⟫ = abs ⟪y, x⟫ := +lemma abs_inner_symm (x y : E) : abs ⟪x, y⟫ = abs ⟪y, x⟫ := by rw [←inner_conj_symm, abs_conj] @[simp] lemma inner_neg_left (x y : E) : ⟪-x, y⟫ = -⟪x, y⟫ := @@ -564,7 +535,7 @@ by rw [←inner_conj_symm, inner_neg_left]; simp only [ring_hom.map_neg, inner_c lemma inner_neg_neg (x y : E) : ⟪-x, -y⟫ = ⟪x, y⟫ := by simp @[simp] lemma inner_self_conj (x : E) : ⟪x, x⟫† = ⟪x, x⟫ := -by rw [is_R_or_C.ext_iff]; exact ⟨by rw [conj_re], by rw [conj_im, inner_self_im_zero, neg_zero]⟩ +by rw [is_R_or_C.ext_iff]; exact ⟨by rw [conj_re], by rw [conj_im, inner_self_im, neg_zero]⟩ lemma inner_sub_left (x y z : E) : ⟪x - y, z⟫ = ⟪x, z⟫ - ⟪y, z⟫ := by { simp [sub_eq_add_neg, inner_add_left] } @@ -637,7 +608,7 @@ begin apply hy', rw is_R_or_C.ext_iff, exact ⟨by simp only [H, zero_re'], - by simp only [inner_self_nonneg_im, add_monoid_hom.map_zero]⟩ }, + by simp only [inner_self_im, add_monoid_hom.map_zero]⟩ }, have h₆ : re ⟪y, y⟫ ≠ 0 := ne_of_gt h₅, have hmain := calc 0 ≤ re ⟪x - T • y, x - T • y⟫ @@ -668,14 +639,9 @@ end /-- Cauchy–Schwarz inequality for real inner products. -/ lemma real_inner_mul_inner_self_le (x y : F) : ⟪x, y⟫_ℝ * ⟪x, y⟫_ℝ ≤ ⟪x, x⟫_ℝ * ⟪y, y⟫_ℝ := -begin - have h₁ : ⟪y, x⟫_ℝ = ⟪x, y⟫_ℝ := by rw [←inner_conj_symm]; refl, - have h₂ := @inner_mul_inner_self_le ℝ F _ _ _ x y, - dsimp at h₂, - have h₃ := abs_mul_abs_self ⟪x, y⟫_ℝ, - rw [h₁] at h₂, - simpa [h₃] using h₂, -end +calc ⟪x, y⟫_ℝ * ⟪x, y⟫_ℝ ≤ is_R_or_C.abs ⟪x, y⟫_ℝ * is_R_or_C.abs ⟪y, x⟫_ℝ : + by { rw [real_inner_comm y, ← is_R_or_C.abs_mul, ← is_R_or_C.norm_eq_abs], exact le_abs_self _ } +... ≤ ⟪x, x⟫_ℝ * ⟪y, y⟫_ℝ : @inner_mul_inner_self_le ℝ _ _ _ _ x y /-- A family of vectors is linearly independent if they are nonzero and orthogonal. -/ @@ -1004,17 +970,8 @@ by { have h := @norm_add_mul_self ℝ _ _ _ _ x y, simpa using h } /-- Expand the square -/ lemma norm_sub_sq (x y : E) : ‖x - y‖^2 = ‖x‖^2 - 2 * (re ⟪x, y⟫) + ‖y‖^2 := -begin - repeat {rw [sq, ←@inner_self_eq_norm_mul_norm 𝕜]}, - rw [inner_sub_sub_self], - calc - re (⟪x, x⟫ - ⟪x, y⟫ - ⟪y, x⟫ + ⟪y, y⟫) - = re ⟪x, x⟫ - re ⟪x, y⟫ - re ⟪y, x⟫ + re ⟪y, y⟫ : by simp only [map_add, map_sub] - ... = -re ⟪y, x⟫ - re ⟪x, y⟫ + re ⟪x, x⟫ + re ⟪y, y⟫ : by ring - ... = -re (⟪x, y⟫†) - re ⟪x, y⟫ + re ⟪x, x⟫ + re ⟪y, y⟫ : by rw [inner_conj_symm] - ... = -re ⟪x, y⟫ - re ⟪x, y⟫ + re ⟪x, x⟫ + re ⟪y, y⟫ : by rw [conj_re] - ... = re ⟪x, x⟫ - 2*re ⟪x, y⟫ + re ⟪y, y⟫ : by ring -end +by rw [sub_eq_add_neg, @norm_add_sq 𝕜 _ _ _ _ x (-y), norm_neg, inner_neg_right, map_neg, mul_neg, + sub_eq_add_neg] alias norm_sub_sq ← norm_sub_pow_two @@ -1039,7 +996,7 @@ begin have : ‖x‖ * ‖y‖ * (‖x‖ * ‖y‖) = (re ⟪x, x⟫) * (re ⟪y, y⟫), simp only [inner_self_eq_norm_mul_norm], ring, rw this, - conv_lhs { congr, skip, rw [inner_abs_conj_symm] }, + conv_lhs { congr, skip, rw [abs_inner_symm] }, exact inner_mul_inner_self_le _ _ end @@ -1428,15 +1385,8 @@ end norms, has absolute value at most 1. -/ lemma abs_real_inner_div_norm_mul_norm_le_one (x y : F) : absR (⟪x, y⟫_ℝ / (‖x‖ * ‖y‖)) ≤ 1 := begin - rw _root_.abs_div, - by_cases h : 0 = absR (‖x‖ * ‖y‖), - { rw [←h, div_zero], - norm_num }, - { change 0 ≠ absR (‖x‖ * ‖y‖) at h, - rw div_le_iff' (lt_of_le_of_ne (ge_iff_le.mp (_root_.abs_nonneg (‖x‖ * ‖y‖))) h), - convert abs_real_inner_le_norm x y using 1, - rw [_root_.abs_mul, _root_.abs_of_nonneg (norm_nonneg x), _root_.abs_of_nonneg (norm_nonneg y), - mul_one] } + rw [_root_.abs_div, _root_.abs_mul, abs_norm, abs_norm], + exact div_le_one_of_le (abs_real_inner_le_norm x y) (by positivity) end /-- The inner product of a vector with a multiple of itself. -/ @@ -1604,23 +1554,10 @@ a negative multiple of the other. -/ lemma real_inner_div_norm_mul_norm_eq_neg_one_iff (x y : F) : ⟪x, y⟫_ℝ / (‖x‖ * ‖y‖) = -1 ↔ (x ≠ 0 ∧ ∃ (r : ℝ), r < 0 ∧ y = r • x) := begin - split, - { intro h, - have ha := h, - apply_fun absR at ha, - norm_num at ha, - rcases (abs_real_inner_div_norm_mul_norm_eq_one_iff x y).1 ha with ⟨hx, ⟨r, ⟨hr, hy⟩⟩⟩, - use [hx, r], - refine and.intro _ hy, - by_contradiction hrpos, - rw hy at h, - rw real_inner_div_norm_mul_norm_eq_one_of_ne_zero_of_pos_mul hx - (lt_of_le_of_ne (le_of_not_lt hrpos) hr.symm) at h, - norm_num at h }, - { intro h, - rcases h with ⟨hx, ⟨r, ⟨hr, hy⟩⟩⟩, - rw hy, - exact real_inner_div_norm_mul_norm_eq_neg_one_of_ne_zero_of_neg_mul hx hr } + rw [← neg_eq_iff_eq_neg, ← neg_div, ← inner_neg_right, ← norm_neg y, + real_inner_div_norm_mul_norm_eq_one_iff, (@neg_surjective ℝ _).exists], + refine iff.rfl.and (exists_congr $ λ r, _), + rw [neg_pos, neg_smul, neg_inj] end /-- If the inner product of two vectors is equal to the product of their norms (i.e., @@ -1722,14 +1659,11 @@ linear_map.mk_continuous₂ (innerₛₗ 𝕜) 1 begin refine le_antisymm ((innerSL 𝕜 x).op_norm_le_bound (norm_nonneg _) (λ y, norm_inner_le_norm _ _)) _, - cases eq_or_lt_of_le (norm_nonneg x) with h h, - { have : x = 0 := norm_eq_zero.mp (eq.symm h), - simp [this] }, - { refine (mul_le_mul_right h).mp _, - calc ‖x‖ * ‖x‖ = ‖x‖ ^ 2 : by ring - ... = re ⟪x, x⟫ : norm_sq_eq_inner _ - ... ≤ abs ⟪x, x⟫ : re_le_abs _ - ... = ‖⟪x, x⟫‖ : by rw [←is_R_or_C.norm_eq_abs] + rcases eq_or_ne x 0 with (rfl | h), + { simp }, + { refine (mul_le_mul_right (norm_pos_iff.2 h)).mp _, + calc ‖x‖ * ‖x‖ = ‖(⟪x, x⟫ : 𝕜)‖ : by rw [← sq, inner_self_eq_norm_sq_to_K, norm_pow, + norm_of_real, norm_norm] ... ≤ ‖innerSL 𝕜 x‖ * ‖x‖ : (innerSL 𝕜 x).le_op_norm _ } end From 73f96237417835f148a1f7bc1ff55f67119b7166 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 3 May 2023 23:08:00 +0000 Subject: [PATCH 0943/1291] chore(ring_theory/kaehler): extract from `ring_theory/derivation` (#18935) This section of the file needs heavier imports than the first half; and this splits the content nicely in two. The lemmas are moved without modification. One very minor docstring typo is fixed. --- src/ring_theory/derivation.lean | 624 +------------------------------ src/ring_theory/etale.lean | 2 +- src/ring_theory/kaehler.lean | 639 ++++++++++++++++++++++++++++++++ 3 files changed, 642 insertions(+), 623 deletions(-) create mode 100644 src/ring_theory/kaehler.lean diff --git a/src/ring_theory/derivation.lean b/src/ring_theory/derivation.lean index 7860637d85bfe..5d70b42f02def 100644 --- a/src/ring_theory/derivation.lean +++ b/src/ring_theory/derivation.lean @@ -4,10 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Nicolò Cavalleri, Andrew Yang -/ -import ring_theory.adjoin.basic import algebra.lie.of_associative -import ring_theory.ideal.cotangent -import ring_theory.is_tensor_product +import ring_theory.adjoin.basic +import ring_theory.ideal.quotient_operations /-! # Derivations @@ -23,25 +22,10 @@ This file defines derivation. A derivation `D` from the `R`-algebra `A` to the ` - `derivation.lie_algebra`: The `R`-derivations from `A` to `A` form an lie algebra over `R`. - `derivation_to_square_zero_equiv_lift`: The `R`-derivations from `A` into a square-zero ideal `I` of `B` corresponds to the lifts `A →ₐ[R] B` of the map `A →ₐ[R] B ⧸ I`. -- `kaehler_differential`: The module of kaehler differentials. For an `R`-algebra `S`, we provide - the notation `Ω[S⁄R]` for `kaehler_differential R S`. - Note that the slash is `\textfractionsolidus`. -- `kaehler_differential.D`: The derivation into the module of kaehler differentials. -- `kaehler_differential.span_range_derivation`: The image of `D` spans `Ω[S⁄R]` as an `S`-module. -- `kaehler_differential.linear_map_equiv_derivation`: - The isomorphism `Hom_R(Ω[S⁄R], M) ≃ₗ[S] Der_R(S, M)`. -- `kaehler_differential.quot_ker_total_equiv`: An alternative description of `Ω[S⁄R]` as `S` copies - of `S` with kernel (`kaehler_differential.ker_total`) generated by the relations: - 1. `dx + dy = d(x + y)` - 2. `x dy + y dx = d(x * y)` - 3. `dr = 0` for `r ∈ R` -- `kaehler_differential.map`: Given a map between the arrows `R → A` and `S → B`, we have an - `A`-linear map `Ω[A⁄R] → Ω[B⁄S]`. ## Future project - Generalize derivations into bimodules. -- Define a `is_kaehler_differential` predicate. -/ @@ -492,607 +476,3 @@ begin end end to_square_zero - -section kaehler_differential - -open_locale tensor_product - -variables (R S : Type*) [comm_ring R] [comm_ring S] [algebra R S] - -/-- The kernel of the multiplication map `S ⊗[R] S →ₐ[R] S`. -/ -abbreviation kaehler_differential.ideal : ideal (S ⊗[R] S) := -ring_hom.ker (tensor_product.lmul' R : S ⊗[R] S →ₐ[R] S) - -variable {S} - -lemma kaehler_differential.one_smul_sub_smul_one_mem_ideal (a : S) : - (1 : S) ⊗ₜ[R] a - a ⊗ₜ[R] (1 : S) ∈ kaehler_differential.ideal R S := -by simp [ring_hom.mem_ker] - -variables {R} - -variables {M : Type*} [add_comm_group M] [module R M] [module S M] [is_scalar_tower R S M] - -/-- For a `R`-derivation `S → M`, this is the map `S ⊗[R] S →ₗ[S] M` sending `s ⊗ₜ t ↦ s • D t`. -/ -def derivation.tensor_product_to (D : derivation R S M) : S ⊗[R] S →ₗ[S] M := -tensor_product.algebra_tensor_module.lift ((linear_map.lsmul S (S →ₗ[R] M)).flip D.to_linear_map) - -lemma derivation.tensor_product_to_tmul (D : derivation R S M) (s t : S) : - D.tensor_product_to (s ⊗ₜ t) = s • D t := -rfl - -lemma derivation.tensor_product_to_mul (D : derivation R S M) (x y : S ⊗[R] S) : - D.tensor_product_to (x * y) = tensor_product.lmul' R x • D.tensor_product_to y + - tensor_product.lmul' R y • D.tensor_product_to x := -begin - apply tensor_product.induction_on x, - { rw [zero_mul, map_zero, map_zero, zero_smul, smul_zero, add_zero] }, - swap, { rintros, simp only [add_mul, map_add, add_smul, *, smul_add], rw add_add_add_comm }, - intros x₁ x₂, - apply tensor_product.induction_on y, - { rw [mul_zero, map_zero, map_zero, zero_smul, smul_zero, add_zero] }, - swap, { rintros, simp only [mul_add, map_add, add_smul, *, smul_add], rw add_add_add_comm }, - intros x y, - simp only [tensor_product.tmul_mul_tmul, derivation.tensor_product_to, - tensor_product.algebra_tensor_module.lift_apply, tensor_product.lift.tmul', - tensor_product.lmul'_apply_tmul], - dsimp, - rw D.leibniz, - simp only [smul_smul, smul_add, mul_comm (x * y) x₁, mul_right_comm x₁ x₂, ← mul_assoc], -end - -variables (R S) - -/-- The kernel of `S ⊗[R] S →ₐ[R] S` is generated by `1 ⊗ s - s ⊗ 1` as a `S`-module. -/ -lemma kaehler_differential.submodule_span_range_eq_ideal : - submodule.span S (set.range $ λ s : S, (1 : S) ⊗ₜ[R] s - s ⊗ₜ[R] (1 : S)) = - (kaehler_differential.ideal R S).restrict_scalars S := -begin - apply le_antisymm, - { rw submodule.span_le, - rintros _ ⟨s, rfl⟩, - exact kaehler_differential.one_smul_sub_smul_one_mem_ideal _ _ }, - { rintros x (hx : _ = _), - have : x - (tensor_product.lmul' R x) ⊗ₜ[R] (1 : S) = x, - { rw [hx, tensor_product.zero_tmul, sub_zero] }, - rw ← this, - clear this hx, - apply tensor_product.induction_on x; clear x, - { rw [map_zero, tensor_product.zero_tmul, sub_zero], exact zero_mem _ }, - { intros x y, - convert_to x • (1 ⊗ₜ y - y ⊗ₜ 1) ∈ _, - { rw [tensor_product.lmul'_apply_tmul, smul_sub, tensor_product.smul_tmul', - tensor_product.smul_tmul', smul_eq_mul, smul_eq_mul, mul_one] }, - { refine submodule.smul_mem _ x _, - apply submodule.subset_span, - exact set.mem_range_self y } }, - { intros x y hx hy, - rw [map_add, tensor_product.add_tmul, ← sub_add_sub_comm], - exact add_mem hx hy } } -end - -lemma kaehler_differential.span_range_eq_ideal : - ideal.span (set.range $ λ s : S, (1 : S) ⊗ₜ[R] s - s ⊗ₜ[R] (1 : S)) = - kaehler_differential.ideal R S := -begin - apply le_antisymm, - { rw ideal.span_le, - rintros _ ⟨s, rfl⟩, - exact kaehler_differential.one_smul_sub_smul_one_mem_ideal _ _ }, - { change (kaehler_differential.ideal R S).restrict_scalars S ≤ (ideal.span _).restrict_scalars S, - rw [← kaehler_differential.submodule_span_range_eq_ideal, ideal.span], - conv_rhs { rw ← submodule.span_span_of_tower S }, - exact submodule.subset_span } -end - -/-- -The module of Kähler differentials (Kahler differentials, Kaehler differentials). -This is implemented as `I / I ^ 2` with `I` the kernel of the multiplication map `S ⊗[R] S →ₐ[R] S`. -To view elements as a linear combination of the form `s • D s'`, use -`kaehler_differential.tensor_product_to_surjective` and `derivation.tensor_product_to_tmul`. - -We also provide the notation `Ω[S⁄R]` for `kaehler_differential R S`. -Note that the slash is `\textfractionsolidus`. --/ -@[derive [add_comm_group, module (S ⊗[R] S)]] -def kaehler_differential : Type* := (kaehler_differential.ideal R S).cotangent - -notation `Ω[`:100 S `⁄`:0 R `]`:0 := kaehler_differential R S - -instance : nonempty Ω[S⁄R] := ⟨0⟩ - -instance kaehler_differential.module' {R' : Type*} [comm_ring R'] [algebra R' S] : - module R' Ω[S⁄R] := -(module.comp_hom (kaehler_differential.ideal R S).cotangent (algebra_map R' S) : _) - -instance : is_scalar_tower S (S ⊗[R] S) Ω[S⁄R] := -ideal.cotangent.is_scalar_tower _ - -instance kaehler_differential.is_scalar_tower_of_tower {R₁ R₂ : Type*} [comm_ring R₁] [comm_ring R₂] - [algebra R₁ S] [algebra R₂ S] [algebra R₁ R₂] [is_scalar_tower R₁ R₂ S] : - is_scalar_tower R₁ R₂ Ω[S⁄R] := -begin - convert restrict_scalars.is_scalar_tower R₁ R₂ Ω[S⁄R] using 1, - ext x m, - show algebra_map R₁ S x • m = algebra_map R₂ S (algebra_map R₁ R₂ x) • m, - rw ← is_scalar_tower.algebra_map_apply, -end - -instance kaehler_differential.is_scalar_tower' : - is_scalar_tower R (S ⊗[R] S) Ω[S⁄R] := -begin - convert restrict_scalars.is_scalar_tower R (S ⊗[R] S) Ω[S⁄R] using 1, - ext x m, - show algebra_map R S x • m = algebra_map R (S ⊗[R] S) x • m, - simp_rw [is_scalar_tower.algebra_map_apply R S (S ⊗[R] S), is_scalar_tower.algebra_map_smul] -end - -/-- The quotient map `I → Ω[S⁄R]` with `I` being the kernel of `S ⊗[R] S → S`. -/ -def kaehler_differential.from_ideal : kaehler_differential.ideal R S →ₗ[S ⊗[R] S] Ω[S⁄R] := -(kaehler_differential.ideal R S).to_cotangent - -/-- (Implementation) The underlying linear map of the derivation into `Ω[S⁄R]`. -/ -def kaehler_differential.D_linear_map : S →ₗ[R] Ω[S⁄R] := -((kaehler_differential.from_ideal R S).restrict_scalars R).comp - ((tensor_product.include_right.to_linear_map - tensor_product.include_left.to_linear_map : - S →ₗ[R] S ⊗[R] S).cod_restrict ((kaehler_differential.ideal R S).restrict_scalars R) - (kaehler_differential.one_smul_sub_smul_one_mem_ideal R) : _ →ₗ[R] _) - -lemma kaehler_differential.D_linear_map_apply (s : S) : - kaehler_differential.D_linear_map R S s = (kaehler_differential.ideal R S).to_cotangent - ⟨1 ⊗ₜ s - s ⊗ₜ 1, kaehler_differential.one_smul_sub_smul_one_mem_ideal R s⟩ := -rfl - -/-- The universal derivation into `Ω[S⁄R]`. -/ -def kaehler_differential.D : derivation R S Ω[S⁄R] := -{ map_one_eq_zero' := begin - dsimp [kaehler_differential.D_linear_map_apply], - rw [ideal.to_cotangent_eq_zero, subtype.coe_mk, sub_self], - exact zero_mem _ - end, - leibniz' := λ a b, begin - dsimp [kaehler_differential.D_linear_map_apply], - rw [← linear_map.map_smul_of_tower, ← linear_map.map_smul_of_tower, ← map_add, - ideal.to_cotangent_eq, pow_two], - convert submodule.mul_mem_mul (kaehler_differential.one_smul_sub_smul_one_mem_ideal R a : _) - (kaehler_differential.one_smul_sub_smul_one_mem_ideal R b : _) using 1, - simp only [add_subgroup_class.coe_sub, submodule.coe_add, submodule.coe_mk, - tensor_product.tmul_mul_tmul, mul_sub, sub_mul, mul_comm b, - submodule.coe_smul_of_tower, smul_sub, tensor_product.smul_tmul', smul_eq_mul, mul_one], - ring_nf, - end, - ..(kaehler_differential.D_linear_map R S) } - -lemma kaehler_differential.D_apply (s : S) : - kaehler_differential.D R S s = (kaehler_differential.ideal R S).to_cotangent - ⟨1 ⊗ₜ s - s ⊗ₜ 1, kaehler_differential.one_smul_sub_smul_one_mem_ideal R s⟩ := -rfl - -lemma kaehler_differential.span_range_derivation : - submodule.span S (set.range $ kaehler_differential.D R S) = ⊤ := -begin - rw _root_.eq_top_iff, - rintros x -, - obtain ⟨⟨x, hx⟩, rfl⟩ := ideal.to_cotangent_surjective _ x, - have : x ∈ (kaehler_differential.ideal R S).restrict_scalars S := hx, - rw ← kaehler_differential.submodule_span_range_eq_ideal at this, - suffices : ∃ hx, (kaehler_differential.ideal R S).to_cotangent ⟨x, hx⟩ ∈ - submodule.span S (set.range $ kaehler_differential.D R S), - { exact this.some_spec }, - apply submodule.span_induction this, - { rintros _ ⟨x, rfl⟩, - refine ⟨kaehler_differential.one_smul_sub_smul_one_mem_ideal R x, _⟩, - apply submodule.subset_span, - exact ⟨x, kaehler_differential.D_linear_map_apply R S x⟩ }, - { exact ⟨zero_mem _, submodule.zero_mem _⟩ }, - { rintros x y ⟨hx₁, hx₂⟩ ⟨hy₁, hy₂⟩, exact ⟨add_mem hx₁ hy₁, submodule.add_mem _ hx₂ hy₂⟩ }, - { rintros r x ⟨hx₁, hx₂⟩, exact ⟨((kaehler_differential.ideal R S).restrict_scalars S).smul_mem - r hx₁, submodule.smul_mem _ r hx₂⟩ } -end - -variables {R S} - -/-- The linear map from `Ω[S⁄R]`, associated with a derivation. -/ -def derivation.lift_kaehler_differential (D : derivation R S M) : Ω[S⁄R] →ₗ[S] M := -begin - refine ((kaehler_differential.ideal R S • ⊤ : - submodule (S ⊗[R] S) (kaehler_differential.ideal R S)).restrict_scalars S).liftq _ _, - { exact D.tensor_product_to.comp ((kaehler_differential.ideal R S).subtype.restrict_scalars S) }, - { intros x hx, - change _ = _, - apply submodule.smul_induction_on hx; clear hx x, - { rintros x (hx : _ = _) ⟨y, hy : _ = _⟩ -, - dsimp, - rw [derivation.tensor_product_to_mul, hx, hy, zero_smul, zero_smul, zero_add] }, - { intros x y ex ey, rw [map_add, ex, ey, zero_add] } } -end - -lemma derivation.lift_kaehler_differential_apply (D : derivation R S M) (x) : - D.lift_kaehler_differential - ((kaehler_differential.ideal R S).to_cotangent x) = D.tensor_product_to x := -rfl - -lemma derivation.lift_kaehler_differential_comp (D : derivation R S M) : - D.lift_kaehler_differential.comp_der (kaehler_differential.D R S) = D := -begin - ext a, - dsimp [kaehler_differential.D_apply], - refine (D.lift_kaehler_differential_apply _).trans _, - rw [subtype.coe_mk, map_sub, derivation.tensor_product_to_tmul, - derivation.tensor_product_to_tmul, one_smul, D.map_one_eq_zero, smul_zero, sub_zero], -end - -@[simp] lemma derivation.lift_kaehler_differential_comp_D (D' : derivation R S M) (x : S) : - D'.lift_kaehler_differential (kaehler_differential.D R S x) = D' x := -derivation.congr_fun D'.lift_kaehler_differential_comp x - -@[ext] -lemma derivation.lift_kaehler_differential_unique - (f f' : Ω[S⁄R] →ₗ[S] M) - (hf : f.comp_der (kaehler_differential.D R S) = - f'.comp_der (kaehler_differential.D R S)) : - f = f' := -begin - apply linear_map.ext, - intro x, - have : x ∈ submodule.span S (set.range $ kaehler_differential.D R S), - { rw kaehler_differential.span_range_derivation, trivial }, - apply submodule.span_induction this, - { rintros _ ⟨x, rfl⟩, exact congr_arg (λ D : derivation R S M, D x) hf }, - { rw [map_zero, map_zero] }, - { intros x y hx hy, rw [map_add, map_add, hx, hy] }, - { intros a x e, rw [map_smul, map_smul, e] } -end - -variables (R S) - -lemma derivation.lift_kaehler_differential_D : - (kaehler_differential.D R S).lift_kaehler_differential = linear_map.id := -derivation.lift_kaehler_differential_unique _ _ - (kaehler_differential.D R S).lift_kaehler_differential_comp - -variables {R S} - -lemma kaehler_differential.D_tensor_product_to (x : kaehler_differential.ideal R S) : - (kaehler_differential.D R S).tensor_product_to x = - (kaehler_differential.ideal R S).to_cotangent x := -begin - rw [← derivation.lift_kaehler_differential_apply, derivation.lift_kaehler_differential_D], - refl, -end - -variables (R S) - -lemma kaehler_differential.tensor_product_to_surjective : - function.surjective (kaehler_differential.D R S).tensor_product_to := -begin - intro x, obtain ⟨x, rfl⟩ := (kaehler_differential.ideal R S).to_cotangent_surjective x, - exact ⟨x, kaehler_differential.D_tensor_product_to x⟩ -end - -/-- The `S`-linear maps from `Ω[S⁄R]` to `M` are (`S`-linearly) equivalent to `R`-derivations -from `S` to `M`. -/ -def kaehler_differential.linear_map_equiv_derivation : (Ω[S⁄R] →ₗ[S] M) ≃ₗ[S] derivation R S M := -{ inv_fun := derivation.lift_kaehler_differential, - left_inv := λ f, derivation.lift_kaehler_differential_unique _ _ - (derivation.lift_kaehler_differential_comp _), - right_inv := derivation.lift_kaehler_differential_comp, - ..(derivation.llcomp.flip $ kaehler_differential.D R S) } - -/-- The quotient ring of `S ⊗ S ⧸ J ^ 2` by `Ω[S⁄R]` is isomorphic to `S`. -/ -def kaehler_differential.quotient_cotangent_ideal_ring_equiv : - (S ⊗ S ⧸ kaehler_differential.ideal R S ^ 2) ⧸ - (kaehler_differential.ideal R S).cotangent_ideal ≃+* S := -begin - have : function.right_inverse tensor_product.include_left - (↑(tensor_product.lmul' R : S ⊗[R] S →ₐ[R] S) : S ⊗[R] S →+* S), - { intro x, rw [alg_hom.coe_to_ring_hom, ← alg_hom.comp_apply, - tensor_product.lmul'_comp_include_left], refl }, - refine (ideal.quot_cotangent _).trans _, - refine (ideal.quot_equiv_of_eq _).trans (ring_hom.quotient_ker_equiv_of_right_inverse this), - ext, refl, -end - -/-- The quotient ring of `S ⊗ S ⧸ J ^ 2` by `Ω[S⁄R]` is isomorphic to `S` as an `S`-algebra. -/ -def kaehler_differential.quotient_cotangent_ideal : - ((S ⊗ S ⧸ kaehler_differential.ideal R S ^ 2) ⧸ - (kaehler_differential.ideal R S).cotangent_ideal) ≃ₐ[S] S := -{ commutes' := (kaehler_differential.quotient_cotangent_ideal_ring_equiv R S).apply_symm_apply, - ..kaehler_differential.quotient_cotangent_ideal_ring_equiv R S } - -lemma kaehler_differential.End_equiv_aux (f : S →ₐ[R] S ⊗ S ⧸ kaehler_differential.ideal R S ^ 2) : - (ideal.quotient.mkₐ R (kaehler_differential.ideal R S).cotangent_ideal).comp f = - is_scalar_tower.to_alg_hom R S _ ↔ - (tensor_product.lmul' R : S ⊗[R] S →ₐ[R] S).ker_square_lift.comp f = alg_hom.id R S := -begin - rw [alg_hom.ext_iff, alg_hom.ext_iff], - apply forall_congr, - intro x, - have e₁ : (tensor_product.lmul' R : S ⊗[R] S →ₐ[R] S).ker_square_lift (f x) = - kaehler_differential.quotient_cotangent_ideal_ring_equiv R S - (ideal.quotient.mk (kaehler_differential.ideal R S).cotangent_ideal $ f x), - { generalize : f x = y, obtain ⟨y, rfl⟩ := ideal.quotient.mk_surjective y, refl }, - have e₂ : x = kaehler_differential.quotient_cotangent_ideal_ring_equiv - R S (is_scalar_tower.to_alg_hom R S _ x), - { exact (mul_one x).symm }, - split, - { intro e, - exact (e₁.trans (@ring_equiv.congr_arg _ _ _ _ _ _ - (kaehler_differential.quotient_cotangent_ideal_ring_equiv R S) _ _ e)).trans e₂.symm }, - { intro e, apply (kaehler_differential.quotient_cotangent_ideal_ring_equiv R S).injective, - exact e₁.symm.trans (e.trans e₂) } -end - -/-- Derivations into `Ω[S⁄R]` is equivalent to derivations -into `(kaehler_differential.ideal R S).cotangent_ideal` -/ --- This has type --- `derivation R S Ω[ S / R ] ≃ₗ[R] derivation R S (kaehler_differential.ideal R S).cotangent_ideal` --- But lean times-out if this is given explicitly. -noncomputable -def kaehler_differential.End_equiv_derivation' := -@linear_equiv.comp_der R _ _ _ _ Ω[S⁄R] _ _ _ _ _ _ _ _ _ - ((kaehler_differential.ideal R S).cotangent_equiv_ideal.restrict_scalars S) - -/-- (Implementation) An `equiv` version of `kaehler_differential.End_equiv_aux`. -Used in `kaehler_differential.End_equiv`. -/ -def kaehler_differential.End_equiv_aux_equiv : - {f // (ideal.quotient.mkₐ R (kaehler_differential.ideal R S).cotangent_ideal).comp f = - is_scalar_tower.to_alg_hom R S _ } ≃ - { f // (tensor_product.lmul' R : S ⊗[R] S →ₐ[R] S).ker_square_lift.comp f = alg_hom.id R S } := -(equiv.refl _).subtype_equiv (kaehler_differential.End_equiv_aux R S) - -/-- -The endomorphisms of `Ω[S⁄R]` corresponds to sections of the surjection `S ⊗[R] S ⧸ J ^ 2 →ₐ[R] S`, -with `J` being the kernel of the multiplication map `S ⊗[R] S →ₐ[R] S`. --/ -noncomputable -def kaehler_differential.End_equiv : - module.End S Ω[S⁄R] ≃ - { f // (tensor_product.lmul' R : S ⊗[R] S →ₐ[R] S).ker_square_lift.comp f = alg_hom.id R S } := -(kaehler_differential.linear_map_equiv_derivation R S).to_equiv.trans $ - (kaehler_differential.End_equiv_derivation' R S).to_equiv.trans $ - (derivation_to_square_zero_equiv_lift - (kaehler_differential.ideal R S).cotangent_ideal - (kaehler_differential.ideal R S).cotangent_ideal_square).trans $ - kaehler_differential.End_equiv_aux_equiv R S - -section presentation - -open kaehler_differential (D) -open finsupp (single) - -/-- The `S`-submodule of `S →₀ S` (the direct sum of copies of `S` indexed by `S`) generated by -the relations: -1. `dx + dy = d(x + y)` -2. `x dy + y dx = d(x * y)` -3. `dr = 0` for `r ∈ R` -where `db` is the unit in the copy of `S` with index `b`. - -This is the kernel of the surjection `finsupp.total S Ω[S⁄R] S (kaehler_differential.D R S)`. -See `kaehler_differential.ker_total_eq` and `kaehler_differential.total_surjective`. --/ -noncomputable -def kaehler_differential.ker_total : submodule S (S →₀ S) := -submodule.span S - ((set.range (λ (x : S × S), single x.1 1 + single x.2 1 - single (x.1 + x.2) 1)) ∪ - (set.range (λ (x : S × S), single x.2 x.1 + single x.1 x.2 - single (x.1 * x.2) 1)) ∪ - (set.range (λ x : R, single (algebra_map R S x) 1))) - -local notation x `𝖣` y := (kaehler_differential.ker_total R S).mkq (single y x) - -lemma kaehler_differential.ker_total_mkq_single_add (x y z) : - (z 𝖣 (x + y)) = (z 𝖣 x) + (z 𝖣 y) := -begin - rw [← map_add, eq_comm, ← sub_eq_zero, ← map_sub, submodule.mkq_apply, - submodule.quotient.mk_eq_zero], - simp_rw [← finsupp.smul_single_one _ z, ← smul_add, ← smul_sub], - exact submodule.smul_mem _ _ (submodule.subset_span (or.inl $ or.inl $ ⟨⟨_, _⟩, rfl⟩)), -end - -lemma kaehler_differential.ker_total_mkq_single_mul (x y z) : - (z 𝖣 (x * y)) = ((z * x) 𝖣 y) + ((z * y) 𝖣 x) := -begin - rw [← map_add, eq_comm, ← sub_eq_zero, ← map_sub, submodule.mkq_apply, - submodule.quotient.mk_eq_zero], - simp_rw [← finsupp.smul_single_one _ z, ← @smul_eq_mul _ _ z, - ← finsupp.smul_single, ← smul_add, ← smul_sub], - exact submodule.smul_mem _ _ (submodule.subset_span (or.inl $ or.inr $ ⟨⟨_, _⟩, rfl⟩)), -end - -lemma kaehler_differential.ker_total_mkq_single_algebra_map (x y) : - (y 𝖣 (algebra_map R S x)) = 0 := -begin - rw [submodule.mkq_apply, submodule.quotient.mk_eq_zero, ← finsupp.smul_single_one _ y], - exact submodule.smul_mem _ _ (submodule.subset_span (or.inr $ ⟨_, rfl⟩)), -end - -lemma kaehler_differential.ker_total_mkq_single_algebra_map_one (x) : - (x 𝖣 1) = 0 := -begin - rw [← (algebra_map R S).map_one, kaehler_differential.ker_total_mkq_single_algebra_map], -end - -lemma kaehler_differential.ker_total_mkq_single_smul (r : R) (x y) : - (y 𝖣 (r • x)) = r • (y 𝖣 x) := -begin - rw [algebra.smul_def, kaehler_differential.ker_total_mkq_single_mul, - kaehler_differential.ker_total_mkq_single_algebra_map, add_zero, - ← linear_map.map_smul_of_tower, finsupp.smul_single, mul_comm, algebra.smul_def], -end - -/-- The (universal) derivation into `(S →₀ S) ⧸ kaehler_differential.ker_total R S`. -/ -noncomputable -def kaehler_differential.derivation_quot_ker_total : - derivation R S ((S →₀ S) ⧸ kaehler_differential.ker_total R S) := -{ to_fun := λ x, 1 𝖣 x, - map_add' := λ x y, kaehler_differential.ker_total_mkq_single_add _ _ _ _ _, - map_smul' := λ r s, kaehler_differential.ker_total_mkq_single_smul _ _ _ _ _, - map_one_eq_zero' := kaehler_differential.ker_total_mkq_single_algebra_map_one _ _ _, - leibniz' := λ a b, (kaehler_differential.ker_total_mkq_single_mul _ _ _ _ _).trans - (by { simp_rw [← finsupp.smul_single_one _ (1 * _ : S)], dsimp, simp }) } - -lemma kaehler_differential.derivation_quot_ker_total_apply (x) : - kaehler_differential.derivation_quot_ker_total R S x = (1 𝖣 x) := rfl - -lemma kaehler_differential.derivation_quot_ker_total_lift_comp_total : - (kaehler_differential.derivation_quot_ker_total R S).lift_kaehler_differential.comp - (finsupp.total S Ω[S⁄R] S (kaehler_differential.D R S)) = submodule.mkq _ := -begin - apply finsupp.lhom_ext, - intros a b, - conv_rhs { rw [← finsupp.smul_single_one a b, linear_map.map_smul] }, - simp [kaehler_differential.derivation_quot_ker_total_apply], -end - -lemma kaehler_differential.ker_total_eq : - (finsupp.total S Ω[S⁄R] S (kaehler_differential.D R S)).ker = - kaehler_differential.ker_total R S := -begin - apply le_antisymm, - { conv_rhs { rw ← (kaehler_differential.ker_total R S).ker_mkq }, - rw ← kaehler_differential.derivation_quot_ker_total_lift_comp_total, - exact linear_map.ker_le_ker_comp _ _ }, - { rw [kaehler_differential.ker_total, submodule.span_le], - rintros _ ((⟨⟨x, y⟩, rfl⟩|⟨⟨x, y⟩, rfl⟩)|⟨x, rfl⟩); dsimp; simp [linear_map.mem_ker] }, -end - -lemma kaehler_differential.total_surjective : - function.surjective (finsupp.total S Ω[S⁄R] S (kaehler_differential.D R S)) := -begin - rw [← linear_map.range_eq_top, finsupp.range_total, kaehler_differential.span_range_derivation], -end - -/-- `Ω[S⁄R]` is isomorphic to `S` copies of `S` with kernel `kaehler_differential.ker_total`. -/ -@[simps] noncomputable -def kaehler_differential.quot_ker_total_equiv : - ((S →₀ S) ⧸ kaehler_differential.ker_total R S) ≃ₗ[S] Ω[S⁄R] := -{ inv_fun := (kaehler_differential.derivation_quot_ker_total R S).lift_kaehler_differential, - left_inv := begin - intro x, - obtain ⟨x, rfl⟩ := submodule.mkq_surjective _ x, - exact linear_map.congr_fun - (kaehler_differential.derivation_quot_ker_total_lift_comp_total R S : _) x, - end, - right_inv := begin - intro x, - obtain ⟨x, rfl⟩ := kaehler_differential.total_surjective R S x, - erw linear_map.congr_fun - (kaehler_differential.derivation_quot_ker_total_lift_comp_total R S : _) x, - refl - end, - ..(kaehler_differential.ker_total R S).liftq - (finsupp.total S Ω[S⁄R] S (kaehler_differential.D R S)) - (kaehler_differential.ker_total_eq R S).ge } - -lemma kaehler_differential.quot_ker_total_equiv_symm_comp_D : - (kaehler_differential.quot_ker_total_equiv R S).symm.to_linear_map.comp_der - (kaehler_differential.D R S) = kaehler_differential.derivation_quot_ker_total R S := -by convert - (kaehler_differential.derivation_quot_ker_total R S).lift_kaehler_differential_comp using 0 - -variables (A B : Type*) [comm_ring A] [comm_ring B] [algebra R A] [algebra S B] [algebra R B] -variables [algebra A B] [is_scalar_tower R S B] [is_scalar_tower R A B] - --- The map `(A →₀ A) →ₗ[A] (B →₀ B)` -local notation `finsupp_map` := ((finsupp.map_range.linear_map (algebra.of_id A B).to_linear_map) - .comp (finsupp.lmap_domain A A (algebra_map A B))) - -lemma kaehler_differential.ker_total_map (h : function.surjective (algebra_map A B)) : - (kaehler_differential.ker_total R A).map finsupp_map ⊔ - submodule.span A (set.range (λ x : S, single (algebra_map S B x) (1 : B))) = - (kaehler_differential.ker_total S B).restrict_scalars _ := -begin - rw [kaehler_differential.ker_total, submodule.map_span, kaehler_differential.ker_total, - submodule.restrict_scalars_span _ _ h], - simp_rw [set.image_union, submodule.span_union, ← set.image_univ, set.image_image, - set.image_univ, map_sub, map_add], - simp only [linear_map.comp_apply, finsupp.map_range.linear_map_apply, finsupp.map_range_single, - finsupp.lmap_domain_apply, finsupp.map_domain_single, alg_hom.to_linear_map_apply, - algebra.of_id_apply, ← is_scalar_tower.algebra_map_apply, map_one, map_add, map_mul], - simp_rw [sup_assoc, ← (h.prod_map h).range_comp], - congr' 3, - rw [sup_eq_right], - apply submodule.span_mono, - simp_rw is_scalar_tower.algebra_map_apply R S B, - exact set.range_comp_subset_range (algebra_map R S) (λ x, single (algebra_map S B x) (1 : B)) -end - -end presentation - -section exact_sequence - -local attribute [irreducible] kaehler_differential - -/- We have the commutative diagram -A --→ B -↑ ↑ -| | -R --→ S -/ -variables (A B : Type*) [comm_ring A] [comm_ring B] [algebra R A] [algebra R B] -variables [algebra A B] [algebra S B] [is_scalar_tower R A B] [is_scalar_tower R S B] - -variables {R B} - -/-- For a tower `R → A → B` and an `R`-derivation `B → M`, we may compose with `A → B` to obtain an -`R`-derivation `A → M`. -/ -def derivation.comp_algebra_map [module A M] [module B M] [is_scalar_tower A B M] - (d : derivation R B M) : derivation R A M := -{ map_one_eq_zero' := by simp, - leibniz' := λ a b, by simp, - to_linear_map := d.to_linear_map.comp (is_scalar_tower.to_alg_hom R A B).to_linear_map } - -variables (R B) - -/-- The map `Ω[A⁄R] →ₗ[A] Ω[B⁄R]` given a square -A --→ B -↑ ↑ -| | -R --→ S -/ -def kaehler_differential.map : Ω[A⁄R] →ₗ[A] Ω[B⁄S] := -derivation.lift_kaehler_differential - (((kaehler_differential.D S B).restrict_scalars R).comp_algebra_map A) - -lemma kaehler_differential.map_comp_der : - (kaehler_differential.map R S A B).comp_der (kaehler_differential.D R A) = - (((kaehler_differential.D S B).restrict_scalars R).comp_algebra_map A) := -derivation.lift_kaehler_differential_comp _ - -lemma kaehler_differential.map_D (x : A) : - kaehler_differential.map R S A B (kaehler_differential.D R A x) = - kaehler_differential.D S B (algebra_map A B x) := -derivation.congr_fun (kaehler_differential.map_comp_der R S A B) x - -open is_scalar_tower (to_alg_hom) - -lemma kaehler_differential.map_surjective_of_surjective - (h : function.surjective (algebra_map A B)) : - function.surjective (kaehler_differential.map R S A B) := -begin - rw [← linear_map.range_eq_top, _root_.eq_top_iff, ← @submodule.restrict_scalars_top B A, - ← kaehler_differential.span_range_derivation, submodule.restrict_scalars_span _ _ h, - submodule.span_le], - rintros _ ⟨x, rfl⟩, - obtain ⟨y, rfl⟩ := h x, - rw ← kaehler_differential.map_D R S A B, - exact ⟨_, rfl⟩, -end - -/-- The lift of the map `Ω[A⁄R] →ₗ[A] Ω[B⁄R]` to the base change along `A → B`. -This is the first map in the exact sequence `B ⊗[A] Ω[A⁄R] → Ω[B⁄R] → Ω[B⁄A] → 0`. -/ -noncomputable -def kaehler_differential.map_base_change : B ⊗[A] Ω[A⁄R] →ₗ[B] Ω[B⁄R] := -(tensor_product.is_base_change A Ω[A⁄R] B).lift (kaehler_differential.map R R A B) - -@[simp] -lemma kaehler_differential.map_base_change_tmul (x : B) (y : Ω[A⁄R]) : - kaehler_differential.map_base_change R A B (x ⊗ₜ y) = - x • kaehler_differential.map R R A B y := -begin - conv_lhs { rw [← mul_one x, ← smul_eq_mul, ← tensor_product.smul_tmul', linear_map.map_smul] }, - congr' 1, - exact is_base_change.lift_eq _ _ _ -end - -end exact_sequence - -end kaehler_differential diff --git a/src/ring_theory/etale.lean b/src/ring_theory/etale.lean index f3d9a2a562c67..09b4302def2a3 100644 --- a/src/ring_theory/etale.lean +++ b/src/ring_theory/etale.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Andrew Yang -/ import ring_theory.quotient_nilpotent -import ring_theory.derivation +import ring_theory.kaehler /-! diff --git a/src/ring_theory/kaehler.lean b/src/ring_theory/kaehler.lean new file mode 100644 index 0000000000000..8e910ef0c9657 --- /dev/null +++ b/src/ring_theory/kaehler.lean @@ -0,0 +1,639 @@ +/- +Copyright © 2020 Nicolò Cavalleri. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Nicolò Cavalleri, Andrew Yang +-/ + +import ring_theory.derivation +import ring_theory.ideal.cotangent +import ring_theory.is_tensor_product + +/-! +# The module of kaehler differentials + +## Main results + +- `kaehler_differential`: The module of kaehler differentials. For an `R`-algebra `S`, we provide + the notation `Ω[S⁄R]` for `kaehler_differential R S`. + Note that the slash is `\textfractionsolidus`. +- `kaehler_differential.D`: The derivation into the module of kaehler differentials. +- `kaehler_differential.span_range_derivation`: The image of `D` spans `Ω[S⁄R]` as an `S`-module. +- `kaehler_differential.linear_map_equiv_derivation`: + The isomorphism `Hom_R(Ω[S⁄R], M) ≃ₗ[S] Der_R(S, M)`. +- `kaehler_differential.quot_ker_total_equiv`: An alternative description of `Ω[S⁄R]` as `S` copies + of `S` with kernel (`kaehler_differential.ker_total`) generated by the relations: + 1. `dx + dy = d(x + y)` + 2. `x dy + y dx = d(x * y)` + 3. `dr = 0` for `r ∈ R` +- `kaehler_differential.map`: Given a map between the arrows `R → A` and `S → B`, we have an + `A`-linear map `Ω[A⁄R] → Ω[B⁄S]`. + +## Future project + +- Define a `is_kaehler_differential` predicate. +-/ + +section kaehler_differential + +open_locale tensor_product +open algebra + +variables (R S : Type*) [comm_ring R] [comm_ring S] [algebra R S] + +/-- The kernel of the multiplication map `S ⊗[R] S →ₐ[R] S`. -/ +abbreviation kaehler_differential.ideal : ideal (S ⊗[R] S) := +ring_hom.ker (tensor_product.lmul' R : S ⊗[R] S →ₐ[R] S) + +variable {S} + +lemma kaehler_differential.one_smul_sub_smul_one_mem_ideal (a : S) : + (1 : S) ⊗ₜ[R] a - a ⊗ₜ[R] (1 : S) ∈ kaehler_differential.ideal R S := +by simp [ring_hom.mem_ker] + +variables {R} + +variables {M : Type*} [add_comm_group M] [module R M] [module S M] [is_scalar_tower R S M] + +/-- For a `R`-derivation `S → M`, this is the map `S ⊗[R] S →ₗ[S] M` sending `s ⊗ₜ t ↦ s • D t`. -/ +def derivation.tensor_product_to (D : derivation R S M) : S ⊗[R] S →ₗ[S] M := +tensor_product.algebra_tensor_module.lift ((linear_map.lsmul S (S →ₗ[R] M)).flip D.to_linear_map) + +lemma derivation.tensor_product_to_tmul (D : derivation R S M) (s t : S) : + D.tensor_product_to (s ⊗ₜ t) = s • D t := +rfl + +lemma derivation.tensor_product_to_mul (D : derivation R S M) (x y : S ⊗[R] S) : + D.tensor_product_to (x * y) = tensor_product.lmul' R x • D.tensor_product_to y + + tensor_product.lmul' R y • D.tensor_product_to x := +begin + apply tensor_product.induction_on x, + { rw [zero_mul, map_zero, map_zero, zero_smul, smul_zero, add_zero] }, + swap, { rintros, simp only [add_mul, map_add, add_smul, *, smul_add], rw add_add_add_comm }, + intros x₁ x₂, + apply tensor_product.induction_on y, + { rw [mul_zero, map_zero, map_zero, zero_smul, smul_zero, add_zero] }, + swap, { rintros, simp only [mul_add, map_add, add_smul, *, smul_add], rw add_add_add_comm }, + intros x y, + simp only [tensor_product.tmul_mul_tmul, derivation.tensor_product_to, + tensor_product.algebra_tensor_module.lift_apply, tensor_product.lift.tmul', + tensor_product.lmul'_apply_tmul], + dsimp, + rw D.leibniz, + simp only [smul_smul, smul_add, mul_comm (x * y) x₁, mul_right_comm x₁ x₂, ← mul_assoc], +end + +variables (R S) + +/-- The kernel of `S ⊗[R] S →ₐ[R] S` is generated by `1 ⊗ s - s ⊗ 1` as a `S`-module. -/ +lemma kaehler_differential.submodule_span_range_eq_ideal : + submodule.span S (set.range $ λ s : S, (1 : S) ⊗ₜ[R] s - s ⊗ₜ[R] (1 : S)) = + (kaehler_differential.ideal R S).restrict_scalars S := +begin + apply le_antisymm, + { rw submodule.span_le, + rintros _ ⟨s, rfl⟩, + exact kaehler_differential.one_smul_sub_smul_one_mem_ideal _ _ }, + { rintros x (hx : _ = _), + have : x - (tensor_product.lmul' R x) ⊗ₜ[R] (1 : S) = x, + { rw [hx, tensor_product.zero_tmul, sub_zero] }, + rw ← this, + clear this hx, + apply tensor_product.induction_on x; clear x, + { rw [map_zero, tensor_product.zero_tmul, sub_zero], exact zero_mem _ }, + { intros x y, + convert_to x • (1 ⊗ₜ y - y ⊗ₜ 1) ∈ _, + { rw [tensor_product.lmul'_apply_tmul, smul_sub, tensor_product.smul_tmul', + tensor_product.smul_tmul', smul_eq_mul, smul_eq_mul, mul_one] }, + { refine submodule.smul_mem _ x _, + apply submodule.subset_span, + exact set.mem_range_self y } }, + { intros x y hx hy, + rw [map_add, tensor_product.add_tmul, ← sub_add_sub_comm], + exact add_mem hx hy } } +end + +lemma kaehler_differential.span_range_eq_ideal : + ideal.span (set.range $ λ s : S, (1 : S) ⊗ₜ[R] s - s ⊗ₜ[R] (1 : S)) = + kaehler_differential.ideal R S := +begin + apply le_antisymm, + { rw ideal.span_le, + rintros _ ⟨s, rfl⟩, + exact kaehler_differential.one_smul_sub_smul_one_mem_ideal _ _ }, + { change (kaehler_differential.ideal R S).restrict_scalars S ≤ (ideal.span _).restrict_scalars S, + rw [← kaehler_differential.submodule_span_range_eq_ideal, ideal.span], + conv_rhs { rw ← submodule.span_span_of_tower S }, + exact submodule.subset_span } +end + +/-- +The module of Kähler differentials (Kahler differentials, Kaehler differentials). +This is implemented as `I / I ^ 2` with `I` the kernel of the multiplication map `S ⊗[R] S →ₐ[R] S`. +To view elements as a linear combination of the form `s • D s'`, use +`kaehler_differential.tensor_product_to_surjective` and `derivation.tensor_product_to_tmul`. + +We also provide the notation `Ω[S⁄R]` for `kaehler_differential R S`. +Note that the slash is `\textfractionsolidus`. +-/ +@[derive [add_comm_group, module (S ⊗[R] S)]] +def kaehler_differential : Type* := (kaehler_differential.ideal R S).cotangent + +notation `Ω[`:100 S `⁄`:0 R `]`:0 := kaehler_differential R S + +instance : nonempty Ω[S⁄R] := ⟨0⟩ + +instance kaehler_differential.module' {R' : Type*} [comm_ring R'] [algebra R' S] : + module R' Ω[S⁄R] := +(module.comp_hom (kaehler_differential.ideal R S).cotangent (algebra_map R' S) : _) + +instance : is_scalar_tower S (S ⊗[R] S) Ω[S⁄R] := +ideal.cotangent.is_scalar_tower _ + +instance kaehler_differential.is_scalar_tower_of_tower {R₁ R₂ : Type*} [comm_ring R₁] [comm_ring R₂] + [algebra R₁ S] [algebra R₂ S] [algebra R₁ R₂] [is_scalar_tower R₁ R₂ S] : + is_scalar_tower R₁ R₂ Ω[S⁄R] := +begin + convert restrict_scalars.is_scalar_tower R₁ R₂ Ω[S⁄R] using 1, + ext x m, + show algebra_map R₁ S x • m = algebra_map R₂ S (algebra_map R₁ R₂ x) • m, + rw ← is_scalar_tower.algebra_map_apply, +end + +instance kaehler_differential.is_scalar_tower' : + is_scalar_tower R (S ⊗[R] S) Ω[S⁄R] := +begin + convert restrict_scalars.is_scalar_tower R (S ⊗[R] S) Ω[S⁄R] using 1, + ext x m, + show algebra_map R S x • m = algebra_map R (S ⊗[R] S) x • m, + simp_rw [is_scalar_tower.algebra_map_apply R S (S ⊗[R] S), is_scalar_tower.algebra_map_smul] +end + +/-- The quotient map `I → Ω[S⁄R]` with `I` being the kernel of `S ⊗[R] S → S`. -/ +def kaehler_differential.from_ideal : kaehler_differential.ideal R S →ₗ[S ⊗[R] S] Ω[S⁄R] := +(kaehler_differential.ideal R S).to_cotangent + +/-- (Implementation) The underlying linear map of the derivation into `Ω[S⁄R]`. -/ +def kaehler_differential.D_linear_map : S →ₗ[R] Ω[S⁄R] := +((kaehler_differential.from_ideal R S).restrict_scalars R).comp + ((tensor_product.include_right.to_linear_map - tensor_product.include_left.to_linear_map : + S →ₗ[R] S ⊗[R] S).cod_restrict ((kaehler_differential.ideal R S).restrict_scalars R) + (kaehler_differential.one_smul_sub_smul_one_mem_ideal R) : _ →ₗ[R] _) + +lemma kaehler_differential.D_linear_map_apply (s : S) : + kaehler_differential.D_linear_map R S s = (kaehler_differential.ideal R S).to_cotangent + ⟨1 ⊗ₜ s - s ⊗ₜ 1, kaehler_differential.one_smul_sub_smul_one_mem_ideal R s⟩ := +rfl + +/-- The universal derivation into `Ω[S⁄R]`. -/ +def kaehler_differential.D : derivation R S Ω[S⁄R] := +{ map_one_eq_zero' := begin + dsimp [kaehler_differential.D_linear_map_apply], + rw [ideal.to_cotangent_eq_zero, subtype.coe_mk, sub_self], + exact zero_mem _ + end, + leibniz' := λ a b, begin + dsimp [kaehler_differential.D_linear_map_apply], + rw [← linear_map.map_smul_of_tower, ← linear_map.map_smul_of_tower, ← map_add, + ideal.to_cotangent_eq, pow_two], + convert submodule.mul_mem_mul (kaehler_differential.one_smul_sub_smul_one_mem_ideal R a : _) + (kaehler_differential.one_smul_sub_smul_one_mem_ideal R b : _) using 1, + simp only [add_subgroup_class.coe_sub, submodule.coe_add, submodule.coe_mk, + tensor_product.tmul_mul_tmul, mul_sub, sub_mul, mul_comm b, + submodule.coe_smul_of_tower, smul_sub, tensor_product.smul_tmul', smul_eq_mul, mul_one], + ring_nf, + end, + ..(kaehler_differential.D_linear_map R S) } + +lemma kaehler_differential.D_apply (s : S) : + kaehler_differential.D R S s = (kaehler_differential.ideal R S).to_cotangent + ⟨1 ⊗ₜ s - s ⊗ₜ 1, kaehler_differential.one_smul_sub_smul_one_mem_ideal R s⟩ := +rfl + +lemma kaehler_differential.span_range_derivation : + submodule.span S (set.range $ kaehler_differential.D R S) = ⊤ := +begin + rw _root_.eq_top_iff, + rintros x -, + obtain ⟨⟨x, hx⟩, rfl⟩ := ideal.to_cotangent_surjective _ x, + have : x ∈ (kaehler_differential.ideal R S).restrict_scalars S := hx, + rw ← kaehler_differential.submodule_span_range_eq_ideal at this, + suffices : ∃ hx, (kaehler_differential.ideal R S).to_cotangent ⟨x, hx⟩ ∈ + submodule.span S (set.range $ kaehler_differential.D R S), + { exact this.some_spec }, + apply submodule.span_induction this, + { rintros _ ⟨x, rfl⟩, + refine ⟨kaehler_differential.one_smul_sub_smul_one_mem_ideal R x, _⟩, + apply submodule.subset_span, + exact ⟨x, kaehler_differential.D_linear_map_apply R S x⟩ }, + { exact ⟨zero_mem _, submodule.zero_mem _⟩ }, + { rintros x y ⟨hx₁, hx₂⟩ ⟨hy₁, hy₂⟩, exact ⟨add_mem hx₁ hy₁, submodule.add_mem _ hx₂ hy₂⟩ }, + { rintros r x ⟨hx₁, hx₂⟩, exact ⟨((kaehler_differential.ideal R S).restrict_scalars S).smul_mem + r hx₁, submodule.smul_mem _ r hx₂⟩ } +end + +variables {R S} + +/-- The linear map from `Ω[S⁄R]`, associated with a derivation. -/ +def derivation.lift_kaehler_differential (D : derivation R S M) : Ω[S⁄R] →ₗ[S] M := +begin + refine ((kaehler_differential.ideal R S • ⊤ : + submodule (S ⊗[R] S) (kaehler_differential.ideal R S)).restrict_scalars S).liftq _ _, + { exact D.tensor_product_to.comp ((kaehler_differential.ideal R S).subtype.restrict_scalars S) }, + { intros x hx, + change _ = _, + apply submodule.smul_induction_on hx; clear hx x, + { rintros x (hx : _ = _) ⟨y, hy : _ = _⟩ -, + dsimp, + rw [derivation.tensor_product_to_mul, hx, hy, zero_smul, zero_smul, zero_add] }, + { intros x y ex ey, rw [map_add, ex, ey, zero_add] } } +end + +lemma derivation.lift_kaehler_differential_apply (D : derivation R S M) (x) : + D.lift_kaehler_differential + ((kaehler_differential.ideal R S).to_cotangent x) = D.tensor_product_to x := +rfl + +lemma derivation.lift_kaehler_differential_comp (D : derivation R S M) : + D.lift_kaehler_differential.comp_der (kaehler_differential.D R S) = D := +begin + ext a, + dsimp [kaehler_differential.D_apply], + refine (D.lift_kaehler_differential_apply _).trans _, + rw [subtype.coe_mk, map_sub, derivation.tensor_product_to_tmul, + derivation.tensor_product_to_tmul, one_smul, D.map_one_eq_zero, smul_zero, sub_zero], +end + +@[simp] lemma derivation.lift_kaehler_differential_comp_D (D' : derivation R S M) (x : S) : + D'.lift_kaehler_differential (kaehler_differential.D R S x) = D' x := +derivation.congr_fun D'.lift_kaehler_differential_comp x + +@[ext] +lemma derivation.lift_kaehler_differential_unique + (f f' : Ω[S⁄R] →ₗ[S] M) + (hf : f.comp_der (kaehler_differential.D R S) = + f'.comp_der (kaehler_differential.D R S)) : + f = f' := +begin + apply linear_map.ext, + intro x, + have : x ∈ submodule.span S (set.range $ kaehler_differential.D R S), + { rw kaehler_differential.span_range_derivation, trivial }, + apply submodule.span_induction this, + { rintros _ ⟨x, rfl⟩, exact congr_arg (λ D : derivation R S M, D x) hf }, + { rw [map_zero, map_zero] }, + { intros x y hx hy, rw [map_add, map_add, hx, hy] }, + { intros a x e, rw [map_smul, map_smul, e] } +end + +variables (R S) + +lemma derivation.lift_kaehler_differential_D : + (kaehler_differential.D R S).lift_kaehler_differential = linear_map.id := +derivation.lift_kaehler_differential_unique _ _ + (kaehler_differential.D R S).lift_kaehler_differential_comp + +variables {R S} + +lemma kaehler_differential.D_tensor_product_to (x : kaehler_differential.ideal R S) : + (kaehler_differential.D R S).tensor_product_to x = + (kaehler_differential.ideal R S).to_cotangent x := +begin + rw [← derivation.lift_kaehler_differential_apply, derivation.lift_kaehler_differential_D], + refl, +end + +variables (R S) + +lemma kaehler_differential.tensor_product_to_surjective : + function.surjective (kaehler_differential.D R S).tensor_product_to := +begin + intro x, obtain ⟨x, rfl⟩ := (kaehler_differential.ideal R S).to_cotangent_surjective x, + exact ⟨x, kaehler_differential.D_tensor_product_to x⟩ +end + +/-- The `S`-linear maps from `Ω[S⁄R]` to `M` are (`S`-linearly) equivalent to `R`-derivations +from `S` to `M`. -/ +def kaehler_differential.linear_map_equiv_derivation : (Ω[S⁄R] →ₗ[S] M) ≃ₗ[S] derivation R S M := +{ inv_fun := derivation.lift_kaehler_differential, + left_inv := λ f, derivation.lift_kaehler_differential_unique _ _ + (derivation.lift_kaehler_differential_comp _), + right_inv := derivation.lift_kaehler_differential_comp, + ..(derivation.llcomp.flip $ kaehler_differential.D R S) } + +/-- The quotient ring of `S ⊗ S ⧸ J ^ 2` by `Ω[S⁄R]` is isomorphic to `S`. -/ +def kaehler_differential.quotient_cotangent_ideal_ring_equiv : + (S ⊗ S ⧸ kaehler_differential.ideal R S ^ 2) ⧸ + (kaehler_differential.ideal R S).cotangent_ideal ≃+* S := +begin + have : function.right_inverse tensor_product.include_left + (↑(tensor_product.lmul' R : S ⊗[R] S →ₐ[R] S) : S ⊗[R] S →+* S), + { intro x, rw [alg_hom.coe_to_ring_hom, ← alg_hom.comp_apply, + tensor_product.lmul'_comp_include_left], refl }, + refine (ideal.quot_cotangent _).trans _, + refine (ideal.quot_equiv_of_eq _).trans (ring_hom.quotient_ker_equiv_of_right_inverse this), + ext, refl, +end + +/-- The quotient ring of `S ⊗ S ⧸ J ^ 2` by `Ω[S⁄R]` is isomorphic to `S` as an `S`-algebra. -/ +def kaehler_differential.quotient_cotangent_ideal : + ((S ⊗ S ⧸ kaehler_differential.ideal R S ^ 2) ⧸ + (kaehler_differential.ideal R S).cotangent_ideal) ≃ₐ[S] S := +{ commutes' := (kaehler_differential.quotient_cotangent_ideal_ring_equiv R S).apply_symm_apply, + ..kaehler_differential.quotient_cotangent_ideal_ring_equiv R S } + +lemma kaehler_differential.End_equiv_aux (f : S →ₐ[R] S ⊗ S ⧸ kaehler_differential.ideal R S ^ 2) : + (ideal.quotient.mkₐ R (kaehler_differential.ideal R S).cotangent_ideal).comp f = + is_scalar_tower.to_alg_hom R S _ ↔ + (tensor_product.lmul' R : S ⊗[R] S →ₐ[R] S).ker_square_lift.comp f = alg_hom.id R S := +begin + rw [alg_hom.ext_iff, alg_hom.ext_iff], + apply forall_congr, + intro x, + have e₁ : (tensor_product.lmul' R : S ⊗[R] S →ₐ[R] S).ker_square_lift (f x) = + kaehler_differential.quotient_cotangent_ideal_ring_equiv R S + (ideal.quotient.mk (kaehler_differential.ideal R S).cotangent_ideal $ f x), + { generalize : f x = y, obtain ⟨y, rfl⟩ := ideal.quotient.mk_surjective y, refl }, + have e₂ : x = kaehler_differential.quotient_cotangent_ideal_ring_equiv + R S (is_scalar_tower.to_alg_hom R S _ x), + { exact (mul_one x).symm }, + split, + { intro e, + exact (e₁.trans (@ring_equiv.congr_arg _ _ _ _ _ _ + (kaehler_differential.quotient_cotangent_ideal_ring_equiv R S) _ _ e)).trans e₂.symm }, + { intro e, apply (kaehler_differential.quotient_cotangent_ideal_ring_equiv R S).injective, + exact e₁.symm.trans (e.trans e₂) } +end + +/-- Derivations into `Ω[S⁄R]` is equivalent to derivations +into `(kaehler_differential.ideal R S).cotangent_ideal` -/ +-- This has type +-- `derivation R S Ω[S⁄R] ≃ₗ[R] derivation R S (kaehler_differential.ideal R S).cotangent_ideal` +-- But lean times-out if this is given explicitly. +noncomputable +def kaehler_differential.End_equiv_derivation' := +@linear_equiv.comp_der R _ _ _ _ Ω[S⁄R] _ _ _ _ _ _ _ _ _ + ((kaehler_differential.ideal R S).cotangent_equiv_ideal.restrict_scalars S) + +/-- (Implementation) An `equiv` version of `kaehler_differential.End_equiv_aux`. +Used in `kaehler_differential.End_equiv`. -/ +def kaehler_differential.End_equiv_aux_equiv : + {f // (ideal.quotient.mkₐ R (kaehler_differential.ideal R S).cotangent_ideal).comp f = + is_scalar_tower.to_alg_hom R S _ } ≃ + { f // (tensor_product.lmul' R : S ⊗[R] S →ₐ[R] S).ker_square_lift.comp f = alg_hom.id R S } := +(equiv.refl _).subtype_equiv (kaehler_differential.End_equiv_aux R S) + +/-- +The endomorphisms of `Ω[S⁄R]` corresponds to sections of the surjection `S ⊗[R] S ⧸ J ^ 2 →ₐ[R] S`, +with `J` being the kernel of the multiplication map `S ⊗[R] S →ₐ[R] S`. +-/ +noncomputable +def kaehler_differential.End_equiv : + module.End S Ω[S⁄R] ≃ + { f // (tensor_product.lmul' R : S ⊗[R] S →ₐ[R] S).ker_square_lift.comp f = alg_hom.id R S } := +(kaehler_differential.linear_map_equiv_derivation R S).to_equiv.trans $ + (kaehler_differential.End_equiv_derivation' R S).to_equiv.trans $ + (derivation_to_square_zero_equiv_lift + (kaehler_differential.ideal R S).cotangent_ideal + (kaehler_differential.ideal R S).cotangent_ideal_square).trans $ + kaehler_differential.End_equiv_aux_equiv R S + +section presentation + +open kaehler_differential (D) +open finsupp (single) + +/-- The `S`-submodule of `S →₀ S` (the direct sum of copies of `S` indexed by `S`) generated by +the relations: +1. `dx + dy = d(x + y)` +2. `x dy + y dx = d(x * y)` +3. `dr = 0` for `r ∈ R` +where `db` is the unit in the copy of `S` with index `b`. + +This is the kernel of the surjection `finsupp.total S Ω[S⁄R] S (kaehler_differential.D R S)`. +See `kaehler_differential.ker_total_eq` and `kaehler_differential.total_surjective`. +-/ +noncomputable +def kaehler_differential.ker_total : submodule S (S →₀ S) := +submodule.span S + ((set.range (λ (x : S × S), single x.1 1 + single x.2 1 - single (x.1 + x.2) 1)) ∪ + (set.range (λ (x : S × S), single x.2 x.1 + single x.1 x.2 - single (x.1 * x.2) 1)) ∪ + (set.range (λ x : R, single (algebra_map R S x) 1))) + +local notation x `𝖣` y := (kaehler_differential.ker_total R S).mkq (single y x) + +lemma kaehler_differential.ker_total_mkq_single_add (x y z) : + (z 𝖣 (x + y)) = (z 𝖣 x) + (z 𝖣 y) := +begin + rw [← map_add, eq_comm, ← sub_eq_zero, ← map_sub, submodule.mkq_apply, + submodule.quotient.mk_eq_zero], + simp_rw [← finsupp.smul_single_one _ z, ← smul_add, ← smul_sub], + exact submodule.smul_mem _ _ (submodule.subset_span (or.inl $ or.inl $ ⟨⟨_, _⟩, rfl⟩)), +end + +lemma kaehler_differential.ker_total_mkq_single_mul (x y z) : + (z 𝖣 (x * y)) = ((z * x) 𝖣 y) + ((z * y) 𝖣 x) := +begin + rw [← map_add, eq_comm, ← sub_eq_zero, ← map_sub, submodule.mkq_apply, + submodule.quotient.mk_eq_zero], + simp_rw [← finsupp.smul_single_one _ z, ← @smul_eq_mul _ _ z, + ← finsupp.smul_single, ← smul_add, ← smul_sub], + exact submodule.smul_mem _ _ (submodule.subset_span (or.inl $ or.inr $ ⟨⟨_, _⟩, rfl⟩)), +end + +lemma kaehler_differential.ker_total_mkq_single_algebra_map (x y) : + (y 𝖣 (algebra_map R S x)) = 0 := +begin + rw [submodule.mkq_apply, submodule.quotient.mk_eq_zero, ← finsupp.smul_single_one _ y], + exact submodule.smul_mem _ _ (submodule.subset_span (or.inr $ ⟨_, rfl⟩)), +end + +lemma kaehler_differential.ker_total_mkq_single_algebra_map_one (x) : + (x 𝖣 1) = 0 := +begin + rw [← (algebra_map R S).map_one, kaehler_differential.ker_total_mkq_single_algebra_map], +end + +lemma kaehler_differential.ker_total_mkq_single_smul (r : R) (x y) : + (y 𝖣 (r • x)) = r • (y 𝖣 x) := +begin + rw [algebra.smul_def, kaehler_differential.ker_total_mkq_single_mul, + kaehler_differential.ker_total_mkq_single_algebra_map, add_zero, + ← linear_map.map_smul_of_tower, finsupp.smul_single, mul_comm, algebra.smul_def], +end + +/-- The (universal) derivation into `(S →₀ S) ⧸ kaehler_differential.ker_total R S`. -/ +noncomputable +def kaehler_differential.derivation_quot_ker_total : + derivation R S ((S →₀ S) ⧸ kaehler_differential.ker_total R S) := +{ to_fun := λ x, 1 𝖣 x, + map_add' := λ x y, kaehler_differential.ker_total_mkq_single_add _ _ _ _ _, + map_smul' := λ r s, kaehler_differential.ker_total_mkq_single_smul _ _ _ _ _, + map_one_eq_zero' := kaehler_differential.ker_total_mkq_single_algebra_map_one _ _ _, + leibniz' := λ a b, (kaehler_differential.ker_total_mkq_single_mul _ _ _ _ _).trans + (by { simp_rw [← finsupp.smul_single_one _ (1 * _ : S)], dsimp, simp }) } + +lemma kaehler_differential.derivation_quot_ker_total_apply (x) : + kaehler_differential.derivation_quot_ker_total R S x = (1 𝖣 x) := rfl + +lemma kaehler_differential.derivation_quot_ker_total_lift_comp_total : + (kaehler_differential.derivation_quot_ker_total R S).lift_kaehler_differential.comp + (finsupp.total S Ω[S⁄R] S (kaehler_differential.D R S)) = submodule.mkq _ := +begin + apply finsupp.lhom_ext, + intros a b, + conv_rhs { rw [← finsupp.smul_single_one a b, linear_map.map_smul] }, + simp [kaehler_differential.derivation_quot_ker_total_apply], +end + +lemma kaehler_differential.ker_total_eq : + (finsupp.total S Ω[S⁄R] S (kaehler_differential.D R S)).ker = + kaehler_differential.ker_total R S := +begin + apply le_antisymm, + { conv_rhs { rw ← (kaehler_differential.ker_total R S).ker_mkq }, + rw ← kaehler_differential.derivation_quot_ker_total_lift_comp_total, + exact linear_map.ker_le_ker_comp _ _ }, + { rw [kaehler_differential.ker_total, submodule.span_le], + rintros _ ((⟨⟨x, y⟩, rfl⟩|⟨⟨x, y⟩, rfl⟩)|⟨x, rfl⟩); dsimp; simp [linear_map.mem_ker] }, +end + +lemma kaehler_differential.total_surjective : + function.surjective (finsupp.total S Ω[S⁄R] S (kaehler_differential.D R S)) := +begin + rw [← linear_map.range_eq_top, finsupp.range_total, kaehler_differential.span_range_derivation], +end + +/-- `Ω[S⁄R]` is isomorphic to `S` copies of `S` with kernel `kaehler_differential.ker_total`. -/ +@[simps] noncomputable +def kaehler_differential.quot_ker_total_equiv : + ((S →₀ S) ⧸ kaehler_differential.ker_total R S) ≃ₗ[S] Ω[S⁄R] := +{ inv_fun := (kaehler_differential.derivation_quot_ker_total R S).lift_kaehler_differential, + left_inv := begin + intro x, + obtain ⟨x, rfl⟩ := submodule.mkq_surjective _ x, + exact linear_map.congr_fun + (kaehler_differential.derivation_quot_ker_total_lift_comp_total R S : _) x, + end, + right_inv := begin + intro x, + obtain ⟨x, rfl⟩ := kaehler_differential.total_surjective R S x, + erw linear_map.congr_fun + (kaehler_differential.derivation_quot_ker_total_lift_comp_total R S : _) x, + refl + end, + ..(kaehler_differential.ker_total R S).liftq + (finsupp.total S Ω[S⁄R] S (kaehler_differential.D R S)) + (kaehler_differential.ker_total_eq R S).ge } + +lemma kaehler_differential.quot_ker_total_equiv_symm_comp_D : + (kaehler_differential.quot_ker_total_equiv R S).symm.to_linear_map.comp_der + (kaehler_differential.D R S) = kaehler_differential.derivation_quot_ker_total R S := +by convert + (kaehler_differential.derivation_quot_ker_total R S).lift_kaehler_differential_comp using 0 + +variables (A B : Type*) [comm_ring A] [comm_ring B] [algebra R A] [algebra S B] [algebra R B] +variables [algebra A B] [is_scalar_tower R S B] [is_scalar_tower R A B] + +-- The map `(A →₀ A) →ₗ[A] (B →₀ B)` +local notation `finsupp_map` := ((finsupp.map_range.linear_map (algebra.of_id A B).to_linear_map) + .comp (finsupp.lmap_domain A A (algebra_map A B))) + +lemma kaehler_differential.ker_total_map (h : function.surjective (algebra_map A B)) : + (kaehler_differential.ker_total R A).map finsupp_map ⊔ + submodule.span A (set.range (λ x : S, single (algebra_map S B x) (1 : B))) = + (kaehler_differential.ker_total S B).restrict_scalars _ := +begin + rw [kaehler_differential.ker_total, submodule.map_span, kaehler_differential.ker_total, + submodule.restrict_scalars_span _ _ h], + simp_rw [set.image_union, submodule.span_union, ← set.image_univ, set.image_image, + set.image_univ, map_sub, map_add], + simp only [linear_map.comp_apply, finsupp.map_range.linear_map_apply, finsupp.map_range_single, + finsupp.lmap_domain_apply, finsupp.map_domain_single, alg_hom.to_linear_map_apply, + algebra.of_id_apply, ← is_scalar_tower.algebra_map_apply, map_one, map_add, map_mul], + simp_rw [sup_assoc, ← (h.prod_map h).range_comp], + congr' 3, + rw [sup_eq_right], + apply submodule.span_mono, + simp_rw is_scalar_tower.algebra_map_apply R S B, + exact set.range_comp_subset_range (algebra_map R S) (λ x, single (algebra_map S B x) (1 : B)) +end + +end presentation + +section exact_sequence + +local attribute [irreducible] kaehler_differential + +/- We have the commutative diagram +A --→ B +↑ ↑ +| | +R --→ S -/ +variables (A B : Type*) [comm_ring A] [comm_ring B] [algebra R A] [algebra R B] +variables [algebra A B] [algebra S B] [is_scalar_tower R A B] [is_scalar_tower R S B] + +variables {R B} + +/-- For a tower `R → A → B` and an `R`-derivation `B → M`, we may compose with `A → B` to obtain an +`R`-derivation `A → M`. -/ +def derivation.comp_algebra_map [module A M] [module B M] [is_scalar_tower A B M] + (d : derivation R B M) : derivation R A M := +{ map_one_eq_zero' := by simp, + leibniz' := λ a b, by simp, + to_linear_map := d.to_linear_map.comp (is_scalar_tower.to_alg_hom R A B).to_linear_map } + +variables (R B) + +/-- The map `Ω[A⁄R] →ₗ[A] Ω[B⁄R]` given a square +A --→ B +↑ ↑ +| | +R --→ S -/ +def kaehler_differential.map : Ω[A⁄R] →ₗ[A] Ω[B⁄S] := +derivation.lift_kaehler_differential + (((kaehler_differential.D S B).restrict_scalars R).comp_algebra_map A) + +lemma kaehler_differential.map_comp_der : + (kaehler_differential.map R S A B).comp_der (kaehler_differential.D R A) = + (((kaehler_differential.D S B).restrict_scalars R).comp_algebra_map A) := +derivation.lift_kaehler_differential_comp _ + +lemma kaehler_differential.map_D (x : A) : + kaehler_differential.map R S A B (kaehler_differential.D R A x) = + kaehler_differential.D S B (algebra_map A B x) := +derivation.congr_fun (kaehler_differential.map_comp_der R S A B) x + +open is_scalar_tower (to_alg_hom) + +lemma kaehler_differential.map_surjective_of_surjective + (h : function.surjective (algebra_map A B)) : + function.surjective (kaehler_differential.map R S A B) := +begin + rw [← linear_map.range_eq_top, _root_.eq_top_iff, ← @submodule.restrict_scalars_top B A, + ← kaehler_differential.span_range_derivation, submodule.restrict_scalars_span _ _ h, + submodule.span_le], + rintros _ ⟨x, rfl⟩, + obtain ⟨y, rfl⟩ := h x, + rw ← kaehler_differential.map_D R S A B, + exact ⟨_, rfl⟩, +end + +/-- The lift of the map `Ω[A⁄R] →ₗ[A] Ω[B⁄R]` to the base change along `A → B`. +This is the first map in the exact sequence `B ⊗[A] Ω[A⁄R] → Ω[B⁄R] → Ω[B⁄A] → 0`. -/ +noncomputable +def kaehler_differential.map_base_change : B ⊗[A] Ω[A⁄R] →ₗ[B] Ω[B⁄R] := +(tensor_product.is_base_change A Ω[A⁄R] B).lift (kaehler_differential.map R R A B) + +@[simp] +lemma kaehler_differential.map_base_change_tmul (x : B) (y : Ω[A⁄R]) : + kaehler_differential.map_base_change R A B (x ⊗ₜ y) = + x • kaehler_differential.map R R A B y := +begin + conv_lhs { rw [← mul_one x, ← smul_eq_mul, ← tensor_product.smul_tmul', linear_map.map_smul] }, + congr' 1, + exact is_base_change.lift_eq _ _ _ +end + +end exact_sequence + +end kaehler_differential From f0c8bf9245297a541f468be517f1bde6195105e9 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Thu, 4 May 2023 01:42:12 +0000 Subject: [PATCH 0944/1291] chore(*): removed unneeded imports (#18926) This is another run of https://github.com/leanprover-community/mathlib/pull/17568, scrubbing unnecessary imports. Like last time we only remove genuinely unneeded imports, and leave merely transitively redundant imports alone. (I *still* disagree with the objectors to removing transitively redundant imports Co-authored-by: Scott Morrison Co-authored-by: Eric Wieser --- src/algebra/category/Module/biproducts.lean | 1 - .../computation/approximation_corollaries.lean | 3 ++- src/algebra/direct_limit.lean | 2 +- src/algebra/order/interval.lean | 2 +- src/algebraic_geometry/Spec.lean | 1 - src/algebraic_geometry/elliptic_curve/point.lean | 2 -- src/algebraic_geometry/locally_ringed_space.lean | 1 - src/algebraic_geometry/morphisms/open_immersion.lean | 2 +- src/algebraic_geometry/open_immersion.lean | 1 - src/analysis/bounded_variation.lean | 1 - src/analysis/calculus/bump_function_findim.lean | 1 + src/analysis/complex/upper_half_plane/topology.lean | 1 - src/analysis/constant_speed.lean | 2 -- src/analysis/convex/between.lean | 2 +- src/analysis/convex/cone/proper.lean | 2 +- src/analysis/convex/intrinsic.lean | 1 - src/analysis/fourier/fourier_transform.lean | 3 +-- src/analysis/fourier/riemann_lebesgue_lemma.lean | 2 -- src/analysis/inner_product_space/orthogonal.lean | 6 ------ src/analysis/normed_space/affine_isometry.lean | 2 +- src/analysis/normed_space/compact_operator.lean | 2 +- src/analysis/normed_space/quaternion_exponential.lean | 1 - src/analysis/normed_space/star/spectrum.lean | 1 - src/analysis/special_functions/japanese_bracket.lean | 1 + src/category_theory/abelian/generator.lean | 2 +- src/category_theory/abelian/injective_resolution.lean | 1 - src/category_theory/abelian/projective.lean | 1 - src/category_theory/sites/cover_preserving.lean | 4 ++-- src/data/array/lemmas.lean | 4 ++-- src/data/buffer/basic.lean | 1 + src/data/hash_map.lean | 1 + src/field_theory/galois.lean | 1 - src/field_theory/minpoly/basic.lean | 1 - src/field_theory/minpoly/is_integrally_closed.lean | 1 - src/geometry/euclidean/angle/oriented/basic.lean | 1 - src/geometry/euclidean/angle/oriented/rotation.lean | 1 + src/linear_algebra/bilinear_form.lean | 1 - src/linear_algebra/bilinear_form/tensor_product.lean | 1 - src/linear_algebra/eigenspace.lean | 1 - src/measure_theory/integral/interval_integral.lean | 2 +- src/measure_theory/integral/periodic.lean | 1 + src/measure_theory/measure/complex_lebesgue.lean | 1 + src/measure_theory/measure/haar_of_basis.lean | 1 - src/measure_theory/measure/probability_measure.lean | 1 - src/measure_theory/tactic.lean | 1 + src/number_theory/cyclotomic/basic.lean | 2 -- src/number_theory/diophantine_approximation.lean | 1 - src/number_theory/kummer_dedekind.lean | 1 - src/number_theory/legendre_symbol/mul_character.lean | 2 +- src/number_theory/modular_forms/jacobi_theta.lean | 4 ++-- src/number_theory/number_field/basic.lean | 1 - src/number_theory/padics/padic_integers.lean | 1 - src/number_theory/primes_congruent_one.lean | 1 - src/number_theory/well_approximable.lean | 1 - src/number_theory/zeta_values.lean | 2 -- src/probability/density.lean | 2 +- src/probability/variance.lean | 1 + src/ring_theory/dedekind_domain/adic_valuation.lean | 2 +- src/ring_theory/dedekind_domain/dvr.lean | 6 ++---- src/ring_theory/dedekind_domain/finite_adele_ring.lean | 1 - src/ring_theory/discrete_valuation_ring/tfae.lean | 1 + src/ring_theory/ideal/associated_prime.lean | 5 ++--- src/ring_theory/ideal/norm.lean | 1 - src/ring_theory/nakayama.lean | 2 +- src/ring_theory/polynomial/gauss_lemma.lean | 1 - src/topology/algebra/nonarchimedean/adic_topology.lean | 2 ++ src/topology/instances/add_circle.lean | 1 - src/topology/instances/complex.lean | 4 ++-- 68 files changed, 40 insertions(+), 77 deletions(-) diff --git a/src/algebra/category/Module/biproducts.lean b/src/algebra/category/Module/biproducts.lean index 24be464b445ce..ac7e918c8c678 100644 --- a/src/algebra/category/Module/biproducts.lean +++ b/src/algebra/category/Module/biproducts.lean @@ -5,7 +5,6 @@ Authors: Scott Morrison -/ import algebra.group.pi import category_theory.limits.shapes.biproducts -import algebra.category.Module.limits import algebra.category.Module.abelian import algebra.homology.short_exact.abelian diff --git a/src/algebra/continued_fractions/computation/approximation_corollaries.lean b/src/algebra/continued_fractions/computation/approximation_corollaries.lean index 9c22e4f9aa49b..df71b81f9db55 100644 --- a/src/algebra/continued_fractions/computation/approximation_corollaries.lean +++ b/src/algebra/continued_fractions/computation/approximation_corollaries.lean @@ -6,7 +6,8 @@ Authors: Kevin Kappelmann import algebra.continued_fractions.computation.approximations import algebra.continued_fractions.convergents_equiv import algebra.order.archimedean -import topology.algebra.order.field +import algebra.algebra.basic +import topology.order.basic /-! # Corollaries From Approximation Lemmas (`algebra.continued_fractions.computation.approximations`) diff --git a/src/algebra/direct_limit.lean b/src/algebra/direct_limit.lean index df4262c66a427..891830bcf9987 100644 --- a/src/algebra/direct_limit.lean +++ b/src/algebra/direct_limit.lean @@ -6,7 +6,7 @@ Authors: Kenny Lau, Chris Hughes import data.finset.order import algebra.direct_sum.module import ring_theory.free_comm_ring -import ring_theory.ideal.quotient_operations +import ring_theory.ideal.quotient /-! # Direct limit of modules, abelian groups, rings, and fields. diff --git a/src/algebra/order/interval.lean b/src/algebra/order/interval.lean index aa63d9e1d7b84..77cdf30d29b8e 100644 --- a/src/algebra/order/interval.lean +++ b/src/algebra/order/interval.lean @@ -15,7 +15,7 @@ import tactic.positivity This file defines arithmetic operations on intervals and prove their correctness. Note that this is full precision operations. The essentials of float operations can be found -in `data.fp.basic`. We hsve not yet integrated these with the rest of the library. +in `data.fp.basic`. We have not yet integrated these with the rest of the library. -/ open function set diff --git a/src/algebraic_geometry/Spec.lean b/src/algebraic_geometry/Spec.lean index 2b6fe31315a2c..8663af3575c76 100644 --- a/src/algebraic_geometry/Spec.lean +++ b/src/algebraic_geometry/Spec.lean @@ -5,7 +5,6 @@ Authors: Scott Morrison, Justus Springer -/ import algebraic_geometry.locally_ringed_space import algebraic_geometry.structure_sheaf -import logic.equiv.transfer_instance import ring_theory.localization.localization_localization import topology.sheaves.sheaf_condition.sites import topology.sheaves.functors diff --git a/src/algebraic_geometry/elliptic_curve/point.lean b/src/algebraic_geometry/elliptic_curve/point.lean index eee73fcbb56f5..59c6aa82ed658 100644 --- a/src/algebraic_geometry/elliptic_curve/point.lean +++ b/src/algebraic_geometry/elliptic_curve/point.lean @@ -5,8 +5,6 @@ Authors: David Kurniadi Angdinata -/ import algebraic_geometry.elliptic_curve.weierstrass -import field_theory.galois -- temporary import to enable point notation -import ring_theory.class_group /-! # The group of nonsingular rational points on a Weierstrass curve over a field diff --git a/src/algebraic_geometry/locally_ringed_space.lean b/src/algebraic_geometry/locally_ringed_space.lean index a1f0c6682c4a1..3e88bfb7a6cdf 100644 --- a/src/algebraic_geometry/locally_ringed_space.lean +++ b/src/algebraic_geometry/locally_ringed_space.lean @@ -6,7 +6,6 @@ Authors: Johan Commelin import algebraic_geometry.ringed_space import algebraic_geometry.stalks -import logic.equiv.transfer_instance /-! # The category of locally ringed spaces diff --git a/src/algebraic_geometry/morphisms/open_immersion.lean b/src/algebraic_geometry/morphisms/open_immersion.lean index 1294196c149bf..62465fea56bb8 100644 --- a/src/algebraic_geometry/morphisms/open_immersion.lean +++ b/src/algebraic_geometry/morphisms/open_immersion.lean @@ -3,8 +3,8 @@ Copyright (c) 2022 Andrew Yang. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Andrew Yang -/ -import algebraic_geometry.morphisms.ring_hom_properties import topology.local_at_target +import algebraic_geometry.morphisms.basic /-! diff --git a/src/algebraic_geometry/open_immersion.lean b/src/algebraic_geometry/open_immersion.lean index 66068eb2190b0..87d243476e371 100644 --- a/src/algebraic_geometry/open_immersion.lean +++ b/src/algebraic_geometry/open_immersion.lean @@ -12,7 +12,6 @@ import algebraic_geometry.Scheme import category_theory.limits.shapes.strict_initial import category_theory.limits.shapes.comm_sq import algebra.category.Ring.instances -import topology.local_at_target /-! # Open immersions of structured spaces diff --git a/src/analysis/bounded_variation.lean b/src/analysis/bounded_variation.lean index b224ea147a178..98b2b765cfdf7 100644 --- a/src/analysis/bounded_variation.lean +++ b/src/analysis/bounded_variation.lean @@ -3,7 +3,6 @@ Copyright (c) 2022 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ -import measure_theory.measure.lebesgue import analysis.calculus.monotone import data.set.function import algebra.group.basic diff --git a/src/analysis/calculus/bump_function_findim.lean b/src/analysis/calculus/bump_function_findim.lean index e44927acf9dd9..453e2bc05af9b 100644 --- a/src/analysis/calculus/bump_function_findim.lean +++ b/src/analysis/calculus/bump_function_findim.lean @@ -6,6 +6,7 @@ Authors: Sébastien Gouëzel import analysis.calculus.series import analysis.convolution import analysis.inner_product_space.euclidean_dist +import measure_theory.measure.haar_lebesgue import data.set.pointwise.support /-! diff --git a/src/analysis/complex/upper_half_plane/topology.lean b/src/analysis/complex/upper_half_plane/topology.lean index 9bbe36ca3bba2..c441e99fcfe93 100644 --- a/src/analysis/complex/upper_half_plane/topology.lean +++ b/src/analysis/complex/upper_half_plane/topology.lean @@ -9,7 +9,6 @@ import analysis.convex.normed import analysis.convex.complex import analysis.complex.re_im_topology import topology.homotopy.contractible -import geometry.manifold.mfderiv import geometry.manifold.cont_mdiff_mfderiv /-! diff --git a/src/analysis/constant_speed.lean b/src/analysis/constant_speed.lean index ea5e9d545fe72..c6ac9e016baf7 100644 --- a/src/analysis/constant_speed.lean +++ b/src/analysis/constant_speed.lean @@ -3,8 +3,6 @@ Copyright (c) 2023 Rémi Bottinelli. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémi Bottinelli -/ -import measure_theory.measure.lebesgue -import analysis.calculus.monotone import data.set.function import analysis.bounded_variation import tactic.swap_var diff --git a/src/analysis/convex/between.lean b/src/analysis/convex/between.lean index 5d192618f9e2c..6c1c1e94a8834 100644 --- a/src/analysis/convex/between.lean +++ b/src/analysis/convex/between.lean @@ -6,8 +6,8 @@ Authors: Joseph Myers import data.set.intervals.group import analysis.convex.segment import linear_algebra.affine_space.finite_dimensional -import linear_algebra.affine_space.midpoint_zero import tactic.field_simp +import algebra.char_p.invertible /-! # Betweenness in affine spaces diff --git a/src/analysis/convex/cone/proper.lean b/src/analysis/convex/cone/proper.lean index 2bc6da0856b11..96ee065db043c 100644 --- a/src/analysis/convex/cone/proper.lean +++ b/src/analysis/convex/cone/proper.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Apurva Nakade All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Apurva Nakade -/ -import analysis.inner_product_space.adjoint +import analysis.convex.cone.basic /-! diff --git a/src/analysis/convex/intrinsic.lean b/src/analysis/convex/intrinsic.lean index 035e2ea3e4686..5d36a26c91ffe 100644 --- a/src/analysis/convex/intrinsic.lean +++ b/src/analysis/convex/intrinsic.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Paul Reichert, Yaël Dillies -/ import analysis.normed_space.add_torsor_bases -import analysis.normed_space.linear_isometry /-! # Intrinsic frontier and interior diff --git a/src/analysis/fourier/fourier_transform.lean b/src/analysis/fourier/fourier_transform.lean index cea8feb24a888..c3dd53ccf025c 100644 --- a/src/analysis/fourier/fourier_transform.lean +++ b/src/analysis/fourier/fourier_transform.lean @@ -4,9 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: David Loeffler -/ -import analysis.special_functions.complex.circle import measure_theory.group.integration -import measure_theory.integral.integral_eq_improper +import measure_theory.measure.haar_of_basis /-! # The Fourier transform diff --git a/src/analysis/fourier/riemann_lebesgue_lemma.lean b/src/analysis/fourier/riemann_lebesgue_lemma.lean index 5813bab5b2938..befd529ada06b 100644 --- a/src/analysis/fourier/riemann_lebesgue_lemma.lean +++ b/src/analysis/fourier/riemann_lebesgue_lemma.lean @@ -5,9 +5,7 @@ Authors: David Loeffler -/ import measure_theory.function.continuous_map_dense -import measure_theory.integral.integral_eq_improper import measure_theory.group.integration -import topology.continuous_function.zero_at_infty import analysis.fourier.fourier_transform import analysis.inner_product_space.dual import topology.metric_space.emetric_paracompact diff --git a/src/analysis/inner_product_space/orthogonal.lean b/src/analysis/inner_product_space/orthogonal.lean index 67955868a091d..e97459a37a164 100644 --- a/src/analysis/inner_product_space/orthogonal.lean +++ b/src/analysis/inner_product_space/orthogonal.lean @@ -3,13 +3,7 @@ Copyright (c) 2019 Zhouhang Zhou. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Zhouhang Zhou, Sébastien Gouëzel, Frédéric Dupuis -/ -import algebra.direct_sum.module -import analysis.complex.basic -import analysis.convex.uniform -import analysis.normed_space.completion -import analysis.normed_space.bounded_linear_maps import linear_algebra.bilinear_form - import analysis.inner_product_space.basic /-! diff --git a/src/analysis/normed_space/affine_isometry.lean b/src/analysis/normed_space/affine_isometry.lean index 0f62bc036be7c..c3e5a9e350546 100644 --- a/src/analysis/normed_space/affine_isometry.lean +++ b/src/analysis/normed_space/affine_isometry.lean @@ -7,7 +7,7 @@ import analysis.normed_space.linear_isometry import analysis.normed.group.add_torsor import analysis.normed_space.basic import linear_algebra.affine_space.restrict -import linear_algebra.affine_space.midpoint_zero +import algebra.char_p.invertible /-! # Affine isometries diff --git a/src/analysis/normed_space/compact_operator.lean b/src/analysis/normed_space/compact_operator.lean index 7b1d943287552..793e39757aa21 100644 --- a/src/analysis/normed_space/compact_operator.lean +++ b/src/analysis/normed_space/compact_operator.lean @@ -3,8 +3,8 @@ Copyright (c) 2022 Anatole Dedecker. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Anatole Dedecker -/ -import analysis.normed_space.operator_norm import analysis.locally_convex.bounded +import topology.algebra.module.strong_topology /-! # Compact operators diff --git a/src/analysis/normed_space/quaternion_exponential.lean b/src/analysis/normed_space/quaternion_exponential.lean index dec605b571e51..4a6786af9568d 100644 --- a/src/analysis/normed_space/quaternion_exponential.lean +++ b/src/analysis/normed_space/quaternion_exponential.lean @@ -5,7 +5,6 @@ Authors: Eric Wieser -/ import analysis.quaternion import analysis.normed_space.exponential -import analysis.inner_product_space.pi_L2 import analysis.special_functions.trigonometric.series /-! diff --git a/src/analysis/normed_space/star/spectrum.lean b/src/analysis/normed_space/star/spectrum.lean index 8994342eb2cb8..c9cf202ad89d5 100644 --- a/src/analysis/normed_space/star/spectrum.lean +++ b/src/analysis/normed_space/star/spectrum.lean @@ -5,7 +5,6 @@ Authors: Jireh Loreaux -/ import analysis.normed_space.star.basic import analysis.normed_space.spectrum -import analysis.normed_space.star.exponential import analysis.special_functions.exponential import algebra.star.star_alg_hom diff --git a/src/analysis/special_functions/japanese_bracket.lean b/src/analysis/special_functions/japanese_bracket.lean index 3ed378923a66a..d287916d88f20 100644 --- a/src/analysis/special_functions/japanese_bracket.lean +++ b/src/analysis/special_functions/japanese_bracket.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 Moritz Doll. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Moritz Doll -/ +import measure_theory.measure.haar_lebesgue import measure_theory.integral.layercake /-! diff --git a/src/category_theory/abelian/generator.lean b/src/category_theory/abelian/generator.lean index 806c76e1f6398..ef01a1539af08 100644 --- a/src/category_theory/abelian/generator.lean +++ b/src/category_theory/abelian/generator.lean @@ -7,7 +7,7 @@ import category_theory.abelian.subobject import category_theory.limits.essentially_small import category_theory.preadditive.injective import category_theory.preadditive.generator -import category_theory.preadditive.yoneda.limits +import category_theory.abelian.opposite /-! # A complete abelian category with enough injectives and a separator has an injective coseparator diff --git a/src/category_theory/abelian/injective_resolution.lean b/src/category_theory/abelian/injective_resolution.lean index b72be7a1ac788..fa70a18e89435 100644 --- a/src/category_theory/abelian/injective_resolution.lean +++ b/src/category_theory/abelian/injective_resolution.lean @@ -5,7 +5,6 @@ Authors: Jujian Zhang, Scott Morrison -/ import algebra.homology.quasi_iso import category_theory.preadditive.injective_resolution -import category_theory.abelian.homology import algebra.homology.homotopy_category /-! diff --git a/src/category_theory/abelian/projective.lean b/src/category_theory/abelian/projective.lean index 0868de8756aaa..68156ef0055fd 100644 --- a/src/category_theory/abelian/projective.lean +++ b/src/category_theory/abelian/projective.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Markus Himmel, Scott Morrison, Jakob von Raumer -/ import algebra.homology.quasi_iso -import category_theory.abelian.homology import category_theory.preadditive.projective_resolution import category_theory.preadditive.yoneda.limits import category_theory.preadditive.yoneda.projective diff --git a/src/category_theory/sites/cover_preserving.lean b/src/category_theory/sites/cover_preserving.lean index 46403554b14e5..2b3133f5f98ab 100644 --- a/src/category_theory/sites/cover_preserving.lean +++ b/src/category_theory/sites/cover_preserving.lean @@ -3,9 +3,9 @@ Copyright (c) 2021 Andrew Yang. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Andrew Yang -/ -import category_theory.sites.limits import category_theory.functor.flat -import category_theory.limits.preserves.filtered +import category_theory.sites.sheaf +import tactic.apply_fun /-! # Cover-preserving functors between sites. diff --git a/src/data/array/lemmas.lean b/src/data/array/lemmas.lean index 2f0f22eff241c..c459b4f2800dc 100644 --- a/src/data/array/lemmas.lean +++ b/src/data/array/lemmas.lean @@ -3,8 +3,8 @@ Copyright (c) 2017 Microsoft Corporation. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura, Mario Carneiro -/ -import control.traversable.equiv -import data.vector.basic +import data.fin.basic +import data.list.basic universes u v w diff --git a/src/data/buffer/basic.lean b/src/data/buffer/basic.lean index fe4df9b0ffeef..cdff3d2644bd8 100644 --- a/src/data/buffer/basic.lean +++ b/src/data/buffer/basic.lean @@ -7,6 +7,7 @@ General utility functions for buffers. -/ import data.array.lemmas import control.traversable.instances +import control.traversable.equiv namespace buffer diff --git a/src/data/hash_map.lean b/src/data/hash_map.lean index e51219833c6c8..e9293d372477a 100644 --- a/src/data/hash_map.lean +++ b/src/data/hash_map.lean @@ -6,6 +6,7 @@ Authors: Leonardo de Moura, Mario Carneiro import data.array.lemmas import data.list.join import data.list.range +import data.list.nodup import data.pnat.defs /-! diff --git a/src/field_theory/galois.lean b/src/field_theory/galois.lean index c760d259fb005..f818b32d57dba 100644 --- a/src/field_theory/galois.lean +++ b/src/field_theory/galois.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Thomas Browning, Patrick Lutz -/ -import field_theory.is_alg_closed.algebraic_closure import field_theory.primitive_element import field_theory.fixed import group_theory.group_action.fixing_subgroup diff --git a/src/field_theory/minpoly/basic.lean b/src/field_theory/minpoly/basic.lean index fd66cfcd422f6..1538c6bad695a 100644 --- a/src/field_theory/minpoly/basic.lean +++ b/src/field_theory/minpoly/basic.lean @@ -3,7 +3,6 @@ Copyright (c) 2019 Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes, Johan Commelin -/ -import data.polynomial.field_division import ring_theory.integral_closure /-! diff --git a/src/field_theory/minpoly/is_integrally_closed.lean b/src/field_theory/minpoly/is_integrally_closed.lean index 8351ee5ba343f..837cc2ad4face 100644 --- a/src/field_theory/minpoly/is_integrally_closed.lean +++ b/src/field_theory/minpoly/is_integrally_closed.lean @@ -3,7 +3,6 @@ Copyright (c) 2019 Riccardo Brasca. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Riccardo Brasca, Paul Lezeau, Junyan Xu -/ -import data.polynomial.field_division import ring_theory.adjoin_root import field_theory.minpoly.field import ring_theory.polynomial.gauss_lemma diff --git a/src/geometry/euclidean/angle/oriented/basic.lean b/src/geometry/euclidean/angle/oriented/basic.lean index 1a588447295da..34957f68fe9b2 100644 --- a/src/geometry/euclidean/angle/oriented/basic.lean +++ b/src/geometry/euclidean/angle/oriented/basic.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Joseph Myers, Heather Macbeth -/ import analysis.inner_product_space.two_dim -import analysis.special_functions.complex.circle import geometry.euclidean.angle.unoriented.basic /-! diff --git a/src/geometry/euclidean/angle/oriented/rotation.lean b/src/geometry/euclidean/angle/oriented/rotation.lean index 0b2e8d1bbb0ee..890777df7fb5f 100644 --- a/src/geometry/euclidean/angle/oriented/rotation.lean +++ b/src/geometry/euclidean/angle/oriented/rotation.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 Joseph Myers. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Joseph Myers, Heather Macbeth -/ +import analysis.special_functions.complex.circle import geometry.euclidean.angle.oriented.basic /-! diff --git a/src/linear_algebra/bilinear_form.lean b/src/linear_algebra/bilinear_form.lean index c87fd45ebb10e..8befd53e09164 100644 --- a/src/linear_algebra/bilinear_form.lean +++ b/src/linear_algebra/bilinear_form.lean @@ -6,7 +6,6 @@ Authors: Andreas Swerdlow, Kexing Ying import linear_algebra.dual import linear_algebra.free_module.finite.matrix -import linear_algebra.matrix.to_lin /-! # Bilinear form diff --git a/src/linear_algebra/bilinear_form/tensor_product.lean b/src/linear_algebra/bilinear_form/tensor_product.lean index fb28fadf3d5d4..97fdc9fcd996f 100644 --- a/src/linear_algebra/bilinear_form/tensor_product.lean +++ b/src/linear_algebra/bilinear_form/tensor_product.lean @@ -5,7 +5,6 @@ Authors: Eric Wieser -/ import linear_algebra.bilinear_form import linear_algebra.tensor_product -import linear_algebra.contraction /-! # The bilinear form on a tensor product diff --git a/src/linear_algebra/eigenspace.lean b/src/linear_algebra/eigenspace.lean index f1551a369b413..1ea0617830ae5 100644 --- a/src/linear_algebra/eigenspace.lean +++ b/src/linear_algebra/eigenspace.lean @@ -6,7 +6,6 @@ Authors: Alexander Bentkamp import algebra.algebra.spectrum import order.hom.basic -import linear_algebra.free_module.finite.basic import linear_algebra.general_linear_group /-! diff --git a/src/measure_theory/integral/interval_integral.lean b/src/measure_theory/integral/interval_integral.lean index 4b7abf014c516..15062741e82dc 100644 --- a/src/measure_theory/integral/interval_integral.lean +++ b/src/measure_theory/integral/interval_integral.lean @@ -5,7 +5,7 @@ Authors: Yury G. Kudryashov, Patrick Massot, Sébastien Gouëzel -/ import data.set.intervals.disjoint import measure_theory.integral.set_integral -import measure_theory.measure.haar_lebesgue +import measure_theory.measure.lebesgue /-! # Integral over an interval diff --git a/src/measure_theory/integral/periodic.lean b/src/measure_theory/integral/periodic.lean index c93969a6b5810..288be8f6d012a 100644 --- a/src/measure_theory/integral/periodic.lean +++ b/src/measure_theory/integral/periodic.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov, Alex Kontorovich, Heather Macbeth -/ +import measure_theory.measure.haar_lebesgue import measure_theory.measure.haar_quotient import measure_theory.integral.interval_integral import topology.algebra.order.floor diff --git a/src/measure_theory/measure/complex_lebesgue.lean b/src/measure_theory/measure/complex_lebesgue.lean index dc00eaa45138e..07066e07eec81 100644 --- a/src/measure_theory/measure/complex_lebesgue.lean +++ b/src/measure_theory/measure/complex_lebesgue.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import measure_theory.measure.lebesgue +import measure_theory.measure.haar_of_basis /-! # Lebesgue measure on `ℂ` diff --git a/src/measure_theory/measure/haar_of_basis.lean b/src/measure_theory/measure/haar_of_basis.lean index ed046f5070061..dfbf34e1b7ff7 100644 --- a/src/measure_theory/measure/haar_of_basis.lean +++ b/src/measure_theory/measure/haar_of_basis.lean @@ -5,7 +5,6 @@ Authors: Sébastien Gouëzel -/ import measure_theory.measure.haar import analysis.inner_product_space.pi_L2 -import measure_theory.constructions.pi /-! # Additive Haar measure constructed from a basis diff --git a/src/measure_theory/measure/probability_measure.lean b/src/measure_theory/measure/probability_measure.lean index b56872b82a18a..8e79c08d4f447 100644 --- a/src/measure_theory/measure/probability_measure.lean +++ b/src/measure_theory/measure/probability_measure.lean @@ -5,7 +5,6 @@ Authors: Kalle Kytölä -/ import measure_theory.measure.finite_measure import measure_theory.integral.average -import probability.conditional_probability /-! # Probability measures diff --git a/src/measure_theory/tactic.lean b/src/measure_theory/tactic.lean index 55bccc36e39fb..8cc6cb66f517f 100644 --- a/src/measure_theory/tactic.lean +++ b/src/measure_theory/tactic.lean @@ -7,6 +7,7 @@ import measure_theory.measure.measure_space_def import tactic.auto_cases import tactic.tidy import tactic.with_local_reducibility + /-! # Tactics for measure theory diff --git a/src/number_theory/cyclotomic/basic.lean b/src/number_theory/cyclotomic/basic.lean index 2d936346d0181..05c071519e7f2 100644 --- a/src/number_theory/cyclotomic/basic.lean +++ b/src/number_theory/cyclotomic/basic.lean @@ -3,10 +3,8 @@ Copyright (c) 2021 Riccardo Brasca. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Riccardo Brasca -/ - import ring_theory.polynomial.cyclotomic.basic import number_theory.number_field.basic -import algebra.char_p.algebra import field_theory.galois /-! diff --git a/src/number_theory/diophantine_approximation.lean b/src/number_theory/diophantine_approximation.lean index b3bf3e84d7cce..0eb722539cf24 100644 --- a/src/number_theory/diophantine_approximation.lean +++ b/src/number_theory/diophantine_approximation.lean @@ -5,7 +5,6 @@ Authors: Michael Geißer, Michael Stoll -/ import tactic.basic import data.real.irrational -import combinatorics.pigeonhole /-! # Diophantine Approximation diff --git a/src/number_theory/kummer_dedekind.lean b/src/number_theory/kummer_dedekind.lean index b15acf6f64974..0ee9f05a5c58d 100644 --- a/src/number_theory/kummer_dedekind.lean +++ b/src/number_theory/kummer_dedekind.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Anne Baanen, Paul Lezeau -/ -import ring_theory.algebra_tower import ring_theory.dedekind_domain.ideal import ring_theory.is_adjoin_root diff --git a/src/number_theory/legendre_symbol/mul_character.lean b/src/number_theory/legendre_symbol/mul_character.lean index 648797dd56d15..e3c827cc2fb1b 100644 --- a/src/number_theory/legendre_symbol/mul_character.lean +++ b/src/number_theory/legendre_symbol/mul_character.lean @@ -5,7 +5,7 @@ Authors: Michael Stoll -/ import algebra.char_p.basic import algebra.euclidean_domain.instances -import algebra.group.conj_finite +import data.fintype.units /-! # Multiplicative characters of finite rings and fields diff --git a/src/number_theory/modular_forms/jacobi_theta.lean b/src/number_theory/modular_forms/jacobi_theta.lean index 3711225ad9523..dd6578c6a37ae 100644 --- a/src/number_theory/modular_forms/jacobi_theta.lean +++ b/src/number_theory/modular_forms/jacobi_theta.lean @@ -3,10 +3,10 @@ Copyright (c) 2023 David Loeffler. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: David Loeffler -/ -import number_theory.modular_forms.basic import analysis.special_functions.gaussian -import analysis.calculus.series import analysis.complex.locally_uniform_limit +import analysis.complex.upper_half_plane.functions_bounded_at_infty +import analysis.complex.upper_half_plane.topology /-! # Jacobi's theta function diff --git a/src/number_theory/number_field/basic.lean b/src/number_theory/number_field/basic.lean index 0344e0afbced0..537628af9ebf9 100644 --- a/src/number_theory/number_field/basic.lean +++ b/src/number_theory/number_field/basic.lean @@ -3,7 +3,6 @@ Copyright (c) 2021 Ashvni Narayanan. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Ashvni Narayanan, Anne Baanen -/ - import algebra.char_p.algebra import ring_theory.dedekind_domain.integral_closure diff --git a/src/number_theory/padics/padic_integers.lean b/src/number_theory/padics/padic_integers.lean index 83af846ac091a..aac1d3115ce20 100644 --- a/src/number_theory/padics/padic_integers.lean +++ b/src/number_theory/padics/padic_integers.lean @@ -5,7 +5,6 @@ Authors: Robert Y. Lewis, Mario Carneiro, Johan Commelin -/ import number_theory.padics.padic_numbers import ring_theory.discrete_valuation_ring.basic -import topology.metric_space.cau_seq_filter /-! # p-adic integers diff --git a/src/number_theory/primes_congruent_one.lean b/src/number_theory/primes_congruent_one.lean index 98cb1db00ec5d..21e0107bb3e66 100644 --- a/src/number_theory/primes_congruent_one.lean +++ b/src/number_theory/primes_congruent_one.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Riccardo Brasca -/ -import data.nat.prime_fin import ring_theory.polynomial.cyclotomic.eval /-! diff --git a/src/number_theory/well_approximable.lean b/src/number_theory/well_approximable.lean index 05e5a4fc84720..93e703e48f198 100644 --- a/src/number_theory/well_approximable.lean +++ b/src/number_theory/well_approximable.lean @@ -5,7 +5,6 @@ Authors: Oliver Nash -/ import dynamics.ergodic.add_circle import measure_theory.covering.liminf_limsup -import data.nat.totient /-! # Well-approximable numbers and Gallagher's ergodic theorem diff --git a/src/number_theory/zeta_values.lean b/src/number_theory/zeta_values.lean index 9f3340f03eb37..bb90b628f09d4 100644 --- a/src/number_theory/zeta_values.lean +++ b/src/number_theory/zeta_values.lean @@ -5,12 +5,10 @@ Authors: David Loeffler -/ import number_theory.bernoulli_polynomials -import analysis.special_functions.integrals import measure_theory.integral.interval_integral import analysis.fourier.add_circle import analysis.p_series - /-! # Critical values of the Riemann zeta function diff --git a/src/probability/density.lean b/src/probability/density.lean index 5545c5925759d..2e622d3848004 100644 --- a/src/probability/density.lean +++ b/src/probability/density.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kexing Ying -/ import measure_theory.decomposition.radon_nikodym -import measure_theory.measure.lebesgue +import measure_theory.measure.haar_of_basis /-! # Probability density function diff --git a/src/probability/variance.lean b/src/probability/variance.lean index ce7e4589021ff..a0b0c08c87d91 100644 --- a/src/probability/variance.lean +++ b/src/probability/variance.lean @@ -5,6 +5,7 @@ Authors: Sébastien Gouëzel, Kexing Ying -/ import probability.notation import probability.integration +import measure_theory.function.l2_space /-! # Variance of random variables diff --git a/src/ring_theory/dedekind_domain/adic_valuation.lean b/src/ring_theory/dedekind_domain/adic_valuation.lean index 2dfb9e9d972a8..86e0bac6f8b44 100644 --- a/src/ring_theory/dedekind_domain/adic_valuation.lean +++ b/src/ring_theory/dedekind_domain/adic_valuation.lean @@ -6,8 +6,8 @@ Authors: María Inés de Frutos-Fernández import ring_theory.dedekind_domain.ideal import ring_theory.valuation.extend_to_localization import ring_theory.valuation.valuation_subring -import ring_theory.polynomial.cyclotomic.basic import topology.algebra.valued_field +import algebra.order.group.type_tags /-! # Adic valuations on Dedekind domains diff --git a/src/ring_theory/dedekind_domain/dvr.lean b/src/ring_theory/dedekind_domain/dvr.lean index 64659de77d46e..62b3fd3e1307b 100644 --- a/src/ring_theory/dedekind_domain/dvr.lean +++ b/src/ring_theory/dedekind_domain/dvr.lean @@ -3,11 +3,9 @@ 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 ring_theory.dedekind_domain.ideal -import ring_theory.discrete_valuation_ring.tfae -import ring_theory.localization.at_prime +import ring_theory.localization.localization_localization import ring_theory.localization.submodule - +import ring_theory.discrete_valuation_ring.tfae /-! # Dedekind domains diff --git a/src/ring_theory/dedekind_domain/finite_adele_ring.lean b/src/ring_theory/dedekind_domain/finite_adele_ring.lean index 48c95752f18df..8c599af57d482 100644 --- a/src/ring_theory/dedekind_domain/finite_adele_ring.lean +++ b/src/ring_theory/dedekind_domain/finite_adele_ring.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: María Inés de Frutos-Fernández -/ import ring_theory.dedekind_domain.adic_valuation -import topology.algebra.uniform_ring /-! diff --git a/src/ring_theory/discrete_valuation_ring/tfae.lean b/src/ring_theory/discrete_valuation_ring/tfae.lean index 0981915865428..bae86b2c29902 100644 --- a/src/ring_theory/discrete_valuation_ring/tfae.lean +++ b/src/ring_theory/discrete_valuation_ring/tfae.lean @@ -7,6 +7,7 @@ import ring_theory.ideal.cotangent import ring_theory.dedekind_domain.basic import ring_theory.valuation.valuation_ring import ring_theory.nakayama + /-! # Equivalent conditions for DVR diff --git a/src/ring_theory/ideal/associated_prime.lean b/src/ring_theory/ideal/associated_prime.lean index c85ea9e010818..db5a3e4f96180 100644 --- a/src/ring_theory/ideal/associated_prime.lean +++ b/src/ring_theory/ideal/associated_prime.lean @@ -5,9 +5,8 @@ Authors: Andrew Yang -/ import linear_algebra.span import ring_theory.ideal.operations -import ring_theory.finiteness -import ring_theory.localization.ideal -import ring_theory.ideal.minimal_prime +import ring_theory.ideal.quotient_operations +import ring_theory.noetherian /-! diff --git a/src/ring_theory/ideal/norm.lean b/src/ring_theory/ideal/norm.lean index 5311fe018ee5b..b0e0cdd6fcd75 100644 --- a/src/ring_theory/ideal/norm.lean +++ b/src/ring_theory/ideal/norm.lean @@ -8,7 +8,6 @@ import algebra.char_p.quotient import data.finsupp.fintype import data.int.absolute_value import data.int.associated -import number_theory.ramification_inertia import linear_algebra.free_module.determinant import linear_algebra.free_module.ideal_quotient import ring_theory.dedekind_domain.pid diff --git a/src/ring_theory/nakayama.lean b/src/ring_theory/nakayama.lean index 58714ac589368..e1ee04629f7c5 100644 --- a/src/ring_theory/nakayama.lean +++ b/src/ring_theory/nakayama.lean @@ -3,8 +3,8 @@ Copyright (c) 2021 Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes -/ -import ring_theory.noetherian import ring_theory.jacobson_ideal + /-! # Nakayama's lemma diff --git a/src/ring_theory/polynomial/gauss_lemma.lean b/src/ring_theory/polynomial/gauss_lemma.lean index 6df229789277c..427d8cca5a60d 100644 --- a/src/ring_theory/polynomial/gauss_lemma.lean +++ b/src/ring_theory/polynomial/gauss_lemma.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Aaron Anderson -/ import ring_theory.int.basic -import field_theory.splitting_field import ring_theory.localization.integral import ring_theory.integrally_closed diff --git a/src/topology/algebra/nonarchimedean/adic_topology.lean b/src/topology/algebra/nonarchimedean/adic_topology.lean index c5382d5f73c57..e72e630fc6767 100644 --- a/src/topology/algebra/nonarchimedean/adic_topology.lean +++ b/src/topology/algebra/nonarchimedean/adic_topology.lean @@ -6,7 +6,9 @@ Authors: Patrick Massot import ring_theory.ideal.operations import topology.algebra.nonarchimedean.bases +import topology.uniform_space.completion import topology.algebra.uniform_ring + /-! # Adic topology diff --git a/src/topology/instances/add_circle.lean b/src/topology/instances/add_circle.lean index 944ff5878b3b9..145643fd5f7a2 100644 --- a/src/topology/instances/add_circle.lean +++ b/src/topology/instances/add_circle.lean @@ -7,7 +7,6 @@ import data.nat.totient import algebra.ring.add_aut import group_theory.divisible import group_theory.order_of_element -import ring_theory.int.basic import algebra.order.floor import algebra.order.to_interval_mod import topology.instances.real diff --git a/src/topology/instances/complex.lean b/src/topology/instances/complex.lean index 5b88f2cf55212..bf183b83e7a88 100644 --- a/src/topology/instances/complex.lean +++ b/src/topology/instances/complex.lean @@ -4,9 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Xavier Roblot -/ -import topology.algebra.uniform_field import analysis.complex.basic -import field_theory.adjoin +import field_theory.intermediate_field +import topology.algebra.uniform_ring /-! # Some results about the topology of ℂ From aa1dbeab548d61bc840a099abae4437046eaaf00 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Thu, 4 May 2023 03:36:49 +0000 Subject: [PATCH 0945/1291] chore(analysis/normed_space/extend): golf, add aux lemmas (#18927) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit * Add `linear_map.extend_to_𝕜'_apply_re`, `linear_map.sq_norm_extend_to_𝕜'_apply`, `continuous_linear_map.norm_extend_to_𝕜`, and `continuous_linear_map.norm_extend_to_𝕜'`. * Rename `norm_bound` to `continuous_linear_map.norm_extend_to_𝕜'_bound`. * Golf, use `namespace`s. --- src/analysis/normed_space/extend.lean | 89 ++++++++++--------- .../normed_space/hahn_banach/extension.lean | 2 +- 2 files changed, 47 insertions(+), 44 deletions(-) diff --git a/src/analysis/normed_space/extend.lean b/src/analysis/normed_space/extend.lean index f0a198fc8ed45..bc01ce94c0d2b 100644 --- a/src/analysis/normed_space/extend.lean +++ b/src/analysis/normed_space/extend.lean @@ -32,14 +32,17 @@ Alternate forms which operate on `[is_scalar_tower ℝ 𝕜 F]` instead are prov -/ open is_R_or_C +open_locale complex_conjugate variables {𝕜 : Type*} [is_R_or_C 𝕜] {F : Type*} [seminormed_add_comm_group F] [normed_space 𝕜 F] -local notation `abs𝕜` := @is_R_or_C.abs 𝕜 _ + +namespace linear_map + +variables [module ℝ F] [is_scalar_tower ℝ 𝕜 F] /-- Extend `fr : F →ₗ[ℝ] ℝ` to `F →ₗ[𝕜] 𝕜` in a way that will also be continuous and have its norm bounded by `‖fr‖` if `fr` is continuous. -/ -noncomputable def linear_map.extend_to_𝕜' - [module ℝ F] [is_scalar_tower ℝ 𝕜 F] (fr : F →ₗ[ℝ] ℝ) : F →ₗ[𝕜] 𝕜 := +noncomputable def extend_to_𝕜' (fr : F →ₗ[ℝ] ℝ) : F →ₗ[𝕜] 𝕜 := begin let fc : F → 𝕜 := λ x, (fr x : 𝕜) - (I : 𝕜) * (fr ((I : 𝕜) • x)), have add : ∀ x y : F, fc (x + y) = fc x + fc y, @@ -73,61 +76,57 @@ begin exact { to_fun := fc, map_add' := add, map_smul' := smul_𝕜 } end -lemma linear_map.extend_to_𝕜'_apply [module ℝ F] [is_scalar_tower ℝ 𝕜 F] - (fr : F →ₗ[ℝ] ℝ) (x : F) : +lemma extend_to_𝕜'_apply (fr : F →ₗ[ℝ] ℝ) (x : F) : fr.extend_to_𝕜' x = (fr x : 𝕜) - (I : 𝕜) * fr ((I : 𝕜) • x) := rfl +@[simp] lemma extend_to_𝕜'_apply_re (fr : F →ₗ[ℝ] ℝ) (x : F) : re (fr.extend_to_𝕜' x : 𝕜) = fr x := +by simp only [extend_to_𝕜'_apply, map_sub, zero_mul, mul_zero, sub_zero] with is_R_or_C_simps + +lemma norm_extend_to_𝕜'_apply_sq (f : F →ₗ[ℝ] ℝ) (x : F) : + ‖(f.extend_to_𝕜' x : 𝕜)‖ ^ 2 = f (conj (f.extend_to_𝕜' x : 𝕜) • x) := +calc ‖(f.extend_to_𝕜' x : 𝕜)‖ ^ 2 = re (conj (f.extend_to_𝕜' x) * f.extend_to_𝕜' x : 𝕜) : + by rw [is_R_or_C.conj_mul_eq_norm_sq_left, norm_sq_eq_def', of_real_re] +... = f (conj (f.extend_to_𝕜' x : 𝕜) • x) : + by rw [← smul_eq_mul, ← map_smul, extend_to_𝕜'_apply_re] + +end linear_map + +namespace continuous_linear_map + +variables [normed_space ℝ F] [is_scalar_tower ℝ 𝕜 F] + /-- The norm of the extension is bounded by `‖fr‖`. -/ -lemma norm_bound [normed_space ℝ F] [is_scalar_tower ℝ 𝕜 F] (fr : F →L[ℝ] ℝ) (x : F) : +lemma norm_extend_to_𝕜'_bound (fr : F →L[ℝ] ℝ) (x : F) : ‖(fr.to_linear_map.extend_to_𝕜' x : 𝕜)‖ ≤ ‖fr‖ * ‖x‖ := begin - let lm : F →ₗ[𝕜] 𝕜 := fr.to_linear_map.extend_to_𝕜', - -- We aim to find a `t : 𝕜` such that - -- * `lm (t • x) = fr (t • x)` (so `lm (t • x) = t * lm x ∈ ℝ`) - -- * `‖lm x‖ = ‖lm (t • x)‖` (so `t.abs` must be 1) - -- If `lm x ≠ 0`, `(lm x)⁻¹` satisfies the first requirement, and after normalizing, it - -- satisfies the second. - -- (If `lm x = 0`, the goal is trivial.) + set lm : F →ₗ[𝕜] 𝕜 := fr.to_linear_map.extend_to_𝕜', classical, by_cases h : lm x = 0, { rw [h, norm_zero], apply mul_nonneg; exact norm_nonneg _ }, - let fx := (lm x)⁻¹, - let t := fx / (abs𝕜 fx : 𝕜), - have ht : abs𝕜 t = 1, by field_simp [abs_of_real, of_real_inv, is_R_or_C.abs_inv, - is_R_or_C.abs_div, is_R_or_C.abs_abs, h], - have h1 : (fr (t • x) : 𝕜) = lm (t • x), - { apply ext, - { simp only [lm, of_real_re, linear_map.extend_to_𝕜'_apply, mul_re, I_re, of_real_im, zero_mul, - add_monoid_hom.map_sub, sub_zero, mul_zero], - refl }, - { symmetry, - calc im (lm (t • x)) - = im (t * lm x) : by rw [lm.map_smul, smul_eq_mul] - ... = im ((lm x)⁻¹ / (abs𝕜 (lm x)⁻¹) * lm x) : rfl - ... = im (1 / (abs𝕜 (lm x)⁻¹ : 𝕜)) : by rw [div_mul_eq_mul_div, inv_mul_cancel h] - ... = 0 : by rw [← of_real_one, ← of_real_div, of_real_im] - ... = im (fr (t • x) : 𝕜) : by rw [of_real_im] } }, - calc ‖lm x‖ = abs𝕜 t * ‖lm x‖ : by rw [ht, one_mul] - ... = ‖t * lm x‖ : by rw [← norm_eq_abs, norm_mul] - ... = ‖lm (t • x)‖ : by rw [←smul_eq_mul, lm.map_smul] - ... = ‖(fr (t • x) : 𝕜)‖ : by rw h1 - ... = ‖fr (t • x)‖ : by rw [norm_eq_abs, abs_of_real, norm_eq_abs, abs_to_real] - ... ≤ ‖fr‖ * ‖t • x‖ : continuous_linear_map.le_op_norm _ _ - ... = ‖fr‖ * (‖t‖ * ‖x‖) : by rw norm_smul - ... ≤ ‖fr‖ * ‖x‖ : by rw [norm_eq_abs, ht, one_mul] + rw [← mul_le_mul_left (norm_pos_iff.2 h), ← sq], + calc ‖lm x‖ ^ 2 = fr (conj (lm x : 𝕜) • x) : fr.to_linear_map.norm_extend_to_𝕜'_apply_sq x + ... ≤ ‖fr (conj (lm x : 𝕜) • x)‖ : le_abs_self _ + ... ≤ ‖fr‖ * ‖conj (lm x : 𝕜) • x‖ : le_op_norm _ _ + ... = ‖(lm x : 𝕜)‖ * (‖fr‖ * ‖x‖) : by rw [norm_smul, norm_conj, mul_left_comm] end /-- Extend `fr : F →L[ℝ] ℝ` to `F →L[𝕜] 𝕜`. -/ -noncomputable def continuous_linear_map.extend_to_𝕜' [normed_space ℝ F] [is_scalar_tower ℝ 𝕜 F] - (fr : F →L[ℝ] ℝ) : - F →L[𝕜] 𝕜 := -linear_map.mk_continuous _ (‖fr‖) (norm_bound _) +noncomputable def extend_to_𝕜' (fr : F →L[ℝ] ℝ) : F →L[𝕜] 𝕜 := +linear_map.mk_continuous _ (‖fr‖) fr.norm_extend_to_𝕜'_bound -lemma continuous_linear_map.extend_to_𝕜'_apply [normed_space ℝ F] [is_scalar_tower ℝ 𝕜 F] - (fr : F →L[ℝ] ℝ) (x : F) : +lemma extend_to_𝕜'_apply (fr : F →L[ℝ] ℝ) (x : F) : fr.extend_to_𝕜' x = (fr x : 𝕜) - (I : 𝕜) * fr ((I : 𝕜) • x) := rfl +@[simp] lemma norm_extend_to_𝕜' (fr : F →L[ℝ] ℝ) : ‖(fr.extend_to_𝕜' : F →L[𝕜] 𝕜)‖ = ‖fr‖ := +le_antisymm (linear_map.mk_continuous_norm_le _ (norm_nonneg _) _) $ + op_norm_le_bound _ (norm_nonneg _) $ λ x, + calc ‖fr x‖ = ‖re (fr.extend_to_𝕜' x : 𝕜)‖ : congr_arg norm (fr.extend_to_𝕜'_apply_re x).symm + ... ≤ ‖(fr.extend_to_𝕜' x : 𝕜)‖ : (abs_re_le_abs _).trans_eq (norm_eq_abs _).symm + ... ≤ ‖(fr.extend_to_𝕜' : F →L[𝕜] 𝕜)‖ * ‖x‖ : le_op_norm _ _ + +end continuous_linear_map + /-- Extend `fr : restrict_scalars ℝ 𝕜 F →ₗ[ℝ] ℝ` to `F →ₗ[𝕜] 𝕜`. -/ noncomputable def linear_map.extend_to_𝕜 (fr : (restrict_scalars ℝ 𝕜 F) →ₗ[ℝ] ℝ) : F →ₗ[𝕜] 𝕜 := fr.extend_to_𝕜' @@ -142,3 +141,7 @@ 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 + +@[simp] lemma continuous_linear_map.norm_extend_to_𝕜 (fr : (restrict_scalars ℝ 𝕜 F) →L[ℝ] ℝ) : + ‖fr.extend_to_𝕜‖ = ‖fr‖ := +fr.norm_extend_to_𝕜' diff --git a/src/analysis/normed_space/hahn_banach/extension.lean b/src/analysis/normed_space/hahn_banach/extension.lean index 74898d0c531ca..80ab1610e970d 100644 --- a/src/analysis/normed_space/hahn_banach/extension.lean +++ b/src/analysis/normed_space/hahn_banach/extension.lean @@ -92,7 +92,7 @@ begin -- And we derive the equality of the norms by bounding on both sides. refine ⟨h, le_antisymm _ _⟩, { calc ‖g.extend_to_𝕜‖ - ≤ ‖g‖ : g.extend_to_𝕜.op_norm_le_bound g.op_norm_nonneg (norm_bound _) + = ‖g‖ : g.norm_extend_to_𝕜 ... = ‖fr‖ : hnormeq ... ≤ ‖re_clm‖ * ‖f‖ : continuous_linear_map.op_norm_comp_le _ _ ... = ‖f‖ : by rw [re_clm_norm, one_mul] }, From fe8d0ff42c3c24d789f491dc2622b6cac3d61564 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Thu, 4 May 2023 07:52:13 +0000 Subject: [PATCH 0946/1291] chore(*): add mathlib4 synchronization comments (#18923) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.continued_fractions.basic` * `analysis.convex.complex` * `analysis.special_functions.trigonometric.chebyshev` * `category_theory.category.Groupoid` * `category_theory.functor.left_derived` * `category_theory.monoidal.free.basic` * `category_theory.monoidal.tor` * `category_theory.sites.left_exact` * `data.complex.determinant` * `data.complex.module` * `data.matrix.rank` * `linear_algebra.affine_space.finite_dimensional` * `linear_algebra.affine_space.matrix` * `linear_algebra.determinant` * `linear_algebra.free_module.determinant` * `measure_theory.measure.outer_measure` * `ring_theory.localization.inv_submonoid` * `topology.algebra.module.determinant` --- src/algebra/continued_fractions/basic.lean | 3 +++ src/analysis/convex/complex.lean | 3 +++ src/analysis/special_functions/trigonometric/chebyshev.lean | 3 +++ src/category_theory/category/Groupoid.lean | 3 +++ src/category_theory/functor/left_derived.lean | 3 +++ src/category_theory/monoidal/free/basic.lean | 3 +++ src/category_theory/monoidal/tor.lean | 3 +++ src/category_theory/sites/left_exact.lean | 3 +++ src/data/complex/determinant.lean | 3 +++ src/data/complex/module.lean | 3 +++ src/data/matrix/rank.lean | 3 +++ src/linear_algebra/affine_space/finite_dimensional.lean | 3 +++ src/linear_algebra/affine_space/matrix.lean | 3 +++ src/linear_algebra/determinant.lean | 3 +++ src/linear_algebra/free_module/determinant.lean | 3 +++ src/measure_theory/measure/outer_measure.lean | 3 +++ src/ring_theory/localization/inv_submonoid.lean | 3 +++ src/topology/algebra/module/determinant.lean | 3 +++ 18 files changed, 54 insertions(+) diff --git a/src/algebra/continued_fractions/basic.lean b/src/algebra/continued_fractions/basic.lean index dfb55dd60930f..8c04dd961e079 100644 --- a/src/algebra/continued_fractions/basic.lean +++ b/src/algebra/continued_fractions/basic.lean @@ -8,6 +8,9 @@ import algebra.field.defs /-! # Basic Definitions/Theorems for Continued Fractions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Summary We define generalised, simple, and regular continued fractions and functions to evaluate their diff --git a/src/analysis/convex/complex.lean b/src/analysis/convex/complex.lean index d22d1c1b96700..58d32f143c96d 100644 --- a/src/analysis/convex/complex.lean +++ b/src/analysis/convex/complex.lean @@ -9,6 +9,9 @@ import data.complex.module /-! # Convexity of half spaces in ℂ +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The open and closed half-spaces in ℂ given by an inequality on either the real or imaginary part are all convex over ℝ. -/ diff --git a/src/analysis/special_functions/trigonometric/chebyshev.lean b/src/analysis/special_functions/trigonometric/chebyshev.lean index ea96d6f5337d2..f35d08a9a9f50 100644 --- a/src/analysis/special_functions/trigonometric/chebyshev.lean +++ b/src/analysis/special_functions/trigonometric/chebyshev.lean @@ -11,6 +11,9 @@ import ring_theory.polynomial.chebyshev /-! # Multiple angle formulas in terms of Chebyshev polynomials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file gives the trigonometric characterizations of Chebyshev polynomials, for both the real (`real.cos`) and complex (`complex.cos`) cosine. -/ diff --git a/src/category_theory/category/Groupoid.lean b/src/category_theory/category/Groupoid.lean index 2d6b843ef1af2..cac379bf9cb22 100644 --- a/src/category_theory/category/Groupoid.lean +++ b/src/category_theory/category/Groupoid.lean @@ -11,6 +11,9 @@ import category_theory.limits.is_limit /-! # Category of groupoids +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains the definition of the category `Groupoid` of all groupoids. In this category objects are groupoids and morphisms are functors between these groupoids. diff --git a/src/category_theory/functor/left_derived.lean b/src/category_theory/functor/left_derived.lean index b8df704a33603..8f8f9baf18c76 100644 --- a/src/category_theory/functor/left_derived.lean +++ b/src/category_theory/functor/left_derived.lean @@ -8,6 +8,9 @@ import category_theory.preadditive.projective_resolution /-! # Left-derived functors +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define the left-derived functors `F.left_derived n : C ⥤ D` for any additive functor `F` out of a category with projective resolutions. diff --git a/src/category_theory/monoidal/free/basic.lean b/src/category_theory/monoidal/free/basic.lean index dd133c7f02b1b..e3f694775af00 100644 --- a/src/category_theory/monoidal/free/basic.lean +++ b/src/category_theory/monoidal/free/basic.lean @@ -8,6 +8,9 @@ import category_theory.monoidal.functor /-! # The free monoidal category over a type +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given a type `C`, the free monoidal category over `C` has as objects formal expressions built from (formal) tensor products of terms of `C` and a formal unit. Its morphisms are compositions and tensor products of identities, unitors and associators. diff --git a/src/category_theory/monoidal/tor.lean b/src/category_theory/monoidal/tor.lean index 446ae7890eff0..9ea38314c8550 100644 --- a/src/category_theory/monoidal/tor.lean +++ b/src/category_theory/monoidal/tor.lean @@ -9,6 +9,9 @@ import category_theory.monoidal.preadditive /-! # Tor, the left-derived functor of tensor product +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define `Tor C n : C ⥤ C ⥤ C`, by left-deriving in the second factor of `(X, Y) ↦ X ⊗ Y`. For now we have almost nothing to say about it! diff --git a/src/category_theory/sites/left_exact.lean b/src/category_theory/sites/left_exact.lean index 74b59a7f4d69a..7288d49ba3824 100644 --- a/src/category_theory/sites/left_exact.lean +++ b/src/category_theory/sites/left_exact.lean @@ -10,6 +10,9 @@ import category_theory.limits.filtered_colimit_commutes_finite_limit /-! # Left exactness of sheafification + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. In this file we show that sheafification commutes with finite limits. -/ diff --git a/src/data/complex/determinant.lean b/src/data/complex/determinant.lean index 931e201d17f4f..331fdff1eb76b 100644 --- a/src/data/complex/determinant.lean +++ b/src/data/complex/determinant.lean @@ -9,6 +9,9 @@ import linear_algebra.determinant /-! # Determinants of maps in the complex numbers as a vector space over `ℝ` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file provides results about the determinants of maps in the complex numbers as a vector space over `ℝ`. diff --git a/src/data/complex/module.lean b/src/data/complex/module.lean index b6a04214d576e..cac447e5921fc 100644 --- a/src/data/complex/module.lean +++ b/src/data/complex/module.lean @@ -12,6 +12,9 @@ import algebra.char_p.invertible /-! # Complex number as a vector space over `ℝ` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains the following instances: * Any `•`-structure (`has_smul`, `mul_action`, `distrib_mul_action`, `module`, `algebra`) on `ℝ` imbues a corresponding structure on `ℂ`. This includes the statement that `ℂ` is an `ℝ` diff --git a/src/data/matrix/rank.lean b/src/data/matrix/rank.lean index 2caf7ce437817..abf18493118dc 100644 --- a/src/data/matrix/rank.lean +++ b/src/data/matrix/rank.lean @@ -13,6 +13,9 @@ import data.complex.module /-! # Rank of matrices +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The rank of a matrix `A` is defined to be the rank of range of the linear map corresponding to `A`. This definition does not depend on the choice of basis, see `matrix.rank_eq_finrank_range_to_lin`. diff --git a/src/linear_algebra/affine_space/finite_dimensional.lean b/src/linear_algebra/affine_space/finite_dimensional.lean index 4b009c060e9c8..07ff73f8731ca 100644 --- a/src/linear_algebra/affine_space/finite_dimensional.lean +++ b/src/linear_algebra/affine_space/finite_dimensional.lean @@ -9,6 +9,9 @@ import linear_algebra.finite_dimensional /-! # Finite-dimensional subspaces of affine spaces. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file provides a few results relating to finite-dimensional subspaces of affine spaces. diff --git a/src/linear_algebra/affine_space/matrix.lean b/src/linear_algebra/affine_space/matrix.lean index 799e8203837cc..fd7eb1c686749 100644 --- a/src/linear_algebra/affine_space/matrix.lean +++ b/src/linear_algebra/affine_space/matrix.lean @@ -9,6 +9,9 @@ import linear_algebra.determinant /-! # Matrix results for barycentric co-ordinates +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Results about the matrix of barycentric co-ordinates for a family of points in an affine space, with respect to some affine basis. -/ diff --git a/src/linear_algebra/determinant.lean b/src/linear_algebra/determinant.lean index c70296ec4e78e..c7a57875c72af 100644 --- a/src/linear_algebra/determinant.lean +++ b/src/linear_algebra/determinant.lean @@ -13,6 +13,9 @@ import linear_algebra.matrix.basis /-! # Determinant of families of vectors +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the determinant of an endomorphism, and of a family of vectors with respect to some basis. For the determinant of a matrix, see the file `linear_algebra.matrix.determinant`. diff --git a/src/linear_algebra/free_module/determinant.lean b/src/linear_algebra/free_module/determinant.lean index 326f70dbf3d04..a9bc68de15a2d 100644 --- a/src/linear_algebra/free_module/determinant.lean +++ b/src/linear_algebra/free_module/determinant.lean @@ -10,6 +10,9 @@ import linear_algebra.free_module.finite.basic /-! # Determinants in free (finite) modules +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Quite a lot of our results on determinants (that you might know in vector spaces) will work for all free (finite) modules over any commutative ring. diff --git a/src/measure_theory/measure/outer_measure.lean b/src/measure_theory/measure/outer_measure.lean index 28fbdf03a3618..e49e132c70227 100644 --- a/src/measure_theory/measure/outer_measure.lean +++ b/src/measure_theory/measure/outer_measure.lean @@ -11,6 +11,9 @@ import data.fin.vec_notation /-! # Outer Measures +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + An outer measure is a function `μ : set α → ℝ≥0∞`, from the powerset of a type to the extended nonnegative real numbers that satisfies the following conditions: 1. `μ ∅ = 0`; diff --git a/src/ring_theory/localization/inv_submonoid.lean b/src/ring_theory/localization/inv_submonoid.lean index d1c056efbe04e..51092aebf7fd4 100644 --- a/src/ring_theory/localization/inv_submonoid.lean +++ b/src/ring_theory/localization/inv_submonoid.lean @@ -11,6 +11,9 @@ import tactic.ring_exp /-! # Submonoid of inverses +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions * `is_localization.inv_submonoid M S` is the submonoid of `S = M⁻¹R` consisting of inverses of diff --git a/src/topology/algebra/module/determinant.lean b/src/topology/algebra/module/determinant.lean index e57201fa2bb0c..4613a6f35a2c7 100644 --- a/src/topology/algebra/module/determinant.lean +++ b/src/topology/algebra/module/determinant.lean @@ -9,6 +9,9 @@ import linear_algebra.determinant /-! # The determinant of a continuous linear map. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ namespace continuous_linear_map From 4e529b03dd62b7b7d13806c3fb974d9d4848910e Mon Sep 17 00:00:00 2001 From: Chris Hughes Date: Thu, 4 May 2023 09:41:23 +0000 Subject: [PATCH 0947/1291] feat(field_theory/ax_grothendieck): Ax-Grothendieck for algebraic extensions of finite fields (#18479) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Chris Hughes <33847686+ChrisHughes24@users.noreply.github.com> Co-authored-by: Yaël Dillies --- src/data/mv_polynomial/basic.lean | 38 +++++++++++++ src/field_theory/ax_grothendieck.lean | 77 +++++++++++++++++++++++++++ 2 files changed, 115 insertions(+) create mode 100644 src/field_theory/ax_grothendieck.lean diff --git a/src/data/mv_polynomial/basic.lean b/src/data/mv_polynomial/basic.lean index 1be5b486ae32f..269d6e8156387 100644 --- a/src/data/mv_polynomial/basic.lean +++ b/src/data/mv_polynomial/basic.lean @@ -1297,6 +1297,44 @@ by { ext, simp only [aeval_X, aeval_tower_X] } end aeval_tower +section eval_mem + +variables {S subS : Type*} [comm_semiring S] [set_like subS S] [subsemiring_class subS S] + +theorem eval₂_mem {f : R →+* S} + {p : mv_polynomial σ R} {s : subS} + (hs : ∀ i ∈ p.support, f (p.coeff i) ∈ s) {v : σ → S} (hv : ∀ i, v i ∈ s) : + mv_polynomial.eval₂ f v p ∈ s := +begin + classical, + replace hs : ∀ i, f (p.coeff i) ∈ s, + { intro i, + by_cases hi : i ∈ p.support, + { exact hs i hi }, + { rw [mv_polynomial.not_mem_support_iff.1 hi, f.map_zero], + exact zero_mem s } }, + induction p using mv_polynomial.induction_on''' with a a b f ha hb0 ih generalizing hs, + { simpa using hs 0 }, + rw [eval₂_add, eval₂_monomial], + refine add_mem (mul_mem _ $ prod_mem $ λ i hi, pow_mem (hv _) _) (ih $ λ i, _), + { simpa only [coeff_add, coeff_monomial, if_pos rfl, + mv_polynomial.not_mem_support_iff.1 ha, add_zero] using hs a }, + have := hs i, + rw [coeff_add, coeff_monomial] at this, + split_ifs at this with h h, + { subst h, + rw [mv_polynomial.not_mem_support_iff.1 ha, map_zero], + exact zero_mem _ }, + { rwa zero_add at this } +end + +theorem eval_mem {p : mv_polynomial σ S} {s : subS} + (hs : ∀ i ∈ p.support, p.coeff i ∈ s) {v : σ → S} (hv : ∀ i, v i ∈ s) : + mv_polynomial.eval v p ∈ s := +eval₂_mem hs hv + +end eval_mem + end comm_semiring end mv_polynomial diff --git a/src/field_theory/ax_grothendieck.lean b/src/field_theory/ax_grothendieck.lean new file mode 100644 index 0000000000000..7cec7336d2e9f --- /dev/null +++ b/src/field_theory/ax_grothendieck.lean @@ -0,0 +1,77 @@ +/- +Copyright (c) 2023 Chris Hughes. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Chris Hughes +-/ +import data.mv_polynomial.basic +import ring_theory.algebraic +import data.fintype.card + +/-! +# Ax-Grothendieck for algebraic extensions of `zmod p` + +This file proves that if `R` is an algebraic extension of a finite field, +then any injective polynomial map `R^n -> R^n` is also surjective. + +This proof is required for the true Ax-Grothendieck theorem, which proves the same result +for any algebraically closed field of characteristic zero. + +## TODO + +The proof of the theorem for characteristic zero is not in mathlib, but it is at +https://github.com/Jlh18/ModelTheoryInLean8 +-/ + +noncomputable theory + +open mv_polynomial finset function + +/-- Any injective polynomial map over an algebraic extension of a finite field is surjective. -/ +lemma ax_grothendieck_of_locally_finite {ι K R : Type*} [field K] [finite K] [comm_ring R] + [finite ι] [algebra K R] (alg : algebra.is_algebraic K R) + (ps : ι → mv_polynomial ι R) + (hinj : injective (λ v i, eval v (ps i))) : + surjective (λ v i, eval v (ps i)) := +begin + have is_int : ∀ x : R, is_integral K x, + from λ x, is_algebraic_iff_is_integral.1 (alg x), + classical, + intros v, + casesI nonempty_fintype ι, + /- `s` is the set of all coefficients of the polynomial, as well as all of + the coordinates of `v`, the point I am trying to find the preimage of. -/ + let s : finset R := finset.bUnion (univ : finset ι) + (λ i, (ps i).support.image (λ x, coeff x (ps i))) + ∪ (univ : finset ι).image v, + have hv : ∀ i, v i ∈ algebra.adjoin K (s : set R), + from λ j, algebra.subset_adjoin + (mem_union_right _ + (mem_image.2 ⟨j, mem_univ _, rfl⟩)), + have hs₁ : ∀ (i : ι) (k : ι →₀ ℕ), k ∈ (ps i).support → + coeff k (ps i) ∈ algebra.adjoin K (s : set R), + from λ i k hk, algebra.subset_adjoin + (mem_union_left _ (mem_bUnion.2 + ⟨i, mem_univ _, mem_image_of_mem _ hk⟩)), + have hs : ∀ i, mv_polynomial.eval v (ps i) ∈ algebra.adjoin K (s : set R), + from λ i, eval_mem (hs₁ _) hv, + letI := is_noetherian_adjoin_finset s (λ x _, is_int x), + letI := module.is_noetherian.finite K (algebra.adjoin K (s : set R)), + letI : finite (algebra.adjoin K (s : set R)) := + finite_dimensional.finite_of_finite + K (algebra.adjoin K (s : set R)), + /- The restriction of the polynomial map, `ps`, to the subalgebra generated by `s` -/ + let res : (ι → algebra.adjoin K (s : set R)) → + (ι → algebra.adjoin K (s : set R)) := + λ x i, ⟨eval (λ j : ι, (x j : R)) (ps i), + eval_mem (hs₁ _) (λ i, (x i).2)⟩, + have hres_inj : injective res, + { intros x y hxy, + ext i, + simp only [res, subtype.ext_iff, funext_iff] at hxy, + exact congr_fun (hinj (funext hxy)) i }, + have hres_surj : surjective res, + from finite.injective_iff_surjective.1 hres_inj, + cases hres_surj (λ i, ⟨v i, hv i⟩) with w hw, + use λ i, w i, + simpa only [res, subtype.ext_iff, funext_iff] using hw, +end From 07992a1d1f7a4176c6d3f160209608be4e198566 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Thu, 4 May 2023 09:41:24 +0000 Subject: [PATCH 0948/1291] chore(analysis/inner_product_space/basic): golf the proof of Cauchy-Schwarz (#18938) ## API changes - Add `inner_product_space.to_core`. - Make `inner_product_space.core` extend `has_inner`. - Rename namespace from `inner_product_space.of_core` to `inner_product_space.core`. - Rename `inner_product_space.of_core.inner_norm_sq_eq_inner_self` to `inner_product_space.core.coe_norm_sq_eq_inner_self`. - Add `inner_product_space.core.norm_inner_symm`. - Add `inner_product_space.core.cauchy_schwarz_aux`, use it to golf the proof of the Cauchy-Schwarz inequality and its versions. - Use norm instead of `is_R_or_C.abs` here and there, the rest will migrate in #18919. - Rename `inner_product_space.of_core.abs_inner_le_norm` to `inner_product_space.core.norm_inner_le_norm`, use norm. - Add `norm_inner_eq_norm_tfae` and `inner_eq_norm_mul_iff_div`. - Rename `abs_inner_div_norm_mul_norm_eq_one_iff` to `norm_inner_div_norm_mul_norm_eq_one_iff`, use norm. - Rename `inner_eq_norm_mul_iff_of_norm_one` to `inner_eq_one_iff_of_norm_one`. Co-authored-by: Johan Commelin --- docs/undergrad.yaml | 2 +- src/analysis/inner_product_space/basic.lean | 423 ++++++++------------ src/analysis/quaternion.lean | 4 +- src/geometry/manifold/instances/sphere.lean | 6 +- src/linear_algebra/matrix/pos_def.lean | 8 +- 5 files changed, 175 insertions(+), 268 deletions(-) diff --git a/docs/undergrad.yaml b/docs/undergrad.yaml index 2b45ce944747e..3008a9394dd97 100644 --- a/docs/undergrad.yaml +++ b/docs/undergrad.yaml @@ -216,7 +216,7 @@ Bilinear and Quadratic Forms Over a Vector Space: dual isomorphism in the euclidean case: 'inner_product_space.to_dual' orthogonal complement: 'submodule.orthogonal' Cauchy-Schwarz inequality: 'inner_mul_inner_self_le' - norm: 'inner_product_space.of_core.to_has_norm' + norm: 'inner_product_space.core.to_has_norm' orthonormal bases: 'maximal_orthonormal_iff_basis_of_finite_dimensional' Endomorphisms: orthogonal group: 'matrix.orthogonal_group' diff --git a/src/analysis/inner_product_space/basic.lean b/src/analysis/inner_product_space/basic.lean index 0c771833488a5..3ef38143a2648 100644 --- a/src/analysis/inner_product_space/basic.lean +++ b/src/analysis/inner_product_space/basic.lean @@ -122,10 +122,8 @@ instance defined on it, otherwise this will create a second non-defeq norm insta /-- A structure requiring that a scalar product is positive definite and symmetric, from which one can construct an `inner_product_space` instance in `inner_product_space.of_core`. -/ @[nolint has_nonempty_instance] -structure inner_product_space.core - (𝕜 : Type*) (F : Type*) - [is_R_or_C 𝕜] [add_comm_group F] [module 𝕜 F] := -(inner : F → F → 𝕜) +structure inner_product_space.core (𝕜 : Type*) (F : Type*) + [is_R_or_C 𝕜] [add_comm_group F] [module 𝕜 F] extends has_inner 𝕜 F := (conj_symm : ∀ x y, conj (inner y x) = inner x y) (nonneg_re : ∀ x, 0 ≤ re (inner x x)) (definite : ∀ x, inner x x = 0 → x = 0) @@ -137,7 +135,18 @@ of the normed space structure that it produces. However, all the instances we wi local to this proof. -/ attribute [class] inner_product_space.core -namespace inner_product_space.of_core +/-- Define `inner_product_space.core` from `inner_product_space`. Defined to reuse lemmas about +`inner_product_space.core` for `inner_product_space`s. Note that the `has_norm` instance provided by +`inner_product_space.core.has_norm` is propositionally but not definitionally equal to the original +norm. -/ +def inner_product_space.to_core [normed_add_comm_group E] [c : inner_product_space 𝕜 E] : + inner_product_space.core 𝕜 E := +{ nonneg_re := λ x, by { rw [← inner_product_space.norm_sq_eq_inner], apply sq_nonneg }, + definite := λ x hx, norm_eq_zero.1 $ pow_eq_zero $ + by rw [inner_product_space.norm_sq_eq_inner x, hx, map_zero], + .. c } + +namespace inner_product_space.core variables [add_comm_group F] [module 𝕜 F] [c : inner_product_space.core 𝕜 F] include c @@ -149,9 +158,11 @@ local notation `absK` := @is_R_or_C.abs 𝕜 _ local notation `ext_iff` := @is_R_or_C.ext_iff 𝕜 _ local postfix `†`:90 := star_ring_end _ -/-- Inner product defined by the `inner_product_space.core` structure. -/ -def to_has_inner : has_inner 𝕜 F := { inner := c.inner } -local attribute [instance] to_has_inner +/-- Inner product defined by the `inner_product_space.core` structure. We can't reuse +`inner_product_space.core.to_has_inner` because it takes `inner_product_space.core` as an explicit +argument. -/ +def to_has_inner' : has_inner 𝕜 F := c.to_has_inner +local attribute [instance] to_has_inner' /-- The norm squared function for `inner_product_space.core` structure. -/ def norm_sq (x : F) := reK ⟪x, x⟫ @@ -171,7 +182,7 @@ c.add_left _ _ _ lemma inner_add_right (x y z : F) : ⟪x, y + z⟫ = ⟪x, y⟫ + ⟪x, z⟫ := by rw [←inner_conj_symm, inner_add_left, ring_hom.map_add]; simp only [inner_conj_symm] -lemma inner_norm_sq_eq_inner_self (x : F) : (norm_sqF x : 𝕜) = ⟪x, x⟫ := +lemma coe_norm_sq_eq_inner_self (x : F) : (norm_sqF x : 𝕜) = ⟪x, x⟫ := begin rw ext_iff, exact ⟨by simp only [of_real_re]; refl, by simp only [inner_self_im, of_real_im]⟩ @@ -211,6 +222,9 @@ by norm_num [ext_iff, inner_self_im] lemma abs_inner_symm (x y : F) : abs ⟪x, y⟫ = abs ⟪y, x⟫ := by rw [←inner_conj_symm, abs_conj] +lemma norm_inner_symm (x y : F) : ‖⟪x, y⟫‖ = ‖⟪y, x⟫‖ := +by rw [is_R_or_C.norm_eq_abs, is_R_or_C.norm_eq_abs, abs_inner_symm] + lemma inner_neg_left (x y : F) : ⟪-x, y⟫ = -⟪x, y⟫ := by { rw [← neg_one_smul 𝕜 x, inner_smul_left], simp } @@ -234,58 +248,36 @@ by simp only [inner_add_left, inner_add_right]; ring lemma inner_sub_sub_self (x y : F) : ⟪x - y, x - y⟫ = ⟪x, x⟫ - ⟪x, y⟫ - ⟪y, x⟫ + ⟪y, y⟫ := by simp only [inner_sub_left, inner_sub_right]; ring +/-- An auxiliary equality useful to prove the **Cauchy–Schwarz inequality**: the square of the norm +of `⟪x, y⟫ • x - ⟪x, x⟫ • y` is equal to `‖x‖ ^ 2 * (‖x‖ ^ 2 * ‖y‖ ^ 2 - ‖⟪x, y⟫‖ ^ 2)`. We use +`inner_product_space.of_core.norm_sq x` etc (defeq to `is_R_or_C.re ⟪x, x⟫`) instead of `‖x‖ ^ 2` +etc to avoid extra rewrites when applying it to an `inner_product_space`. -/ +theorem cauchy_schwarz_aux (x y : F) : + norm_sqF (⟪x, y⟫ • x - ⟪x, x⟫ • y) = + norm_sqF x * (norm_sqF x * norm_sqF y - ‖⟪x, y⟫‖ ^ 2) := +begin + rw [← @of_real_inj 𝕜, coe_norm_sq_eq_inner_self], + simp only [inner_sub_sub_self, inner_smul_left, inner_smul_right, conj_of_real, mul_sub, + ← coe_norm_sq_eq_inner_self x, ← coe_norm_sq_eq_inner_self y], + rw [← mul_assoc, mul_conj, is_R_or_C.conj_mul_eq_norm_sq_left, norm_sq_eq_def', + mul_left_comm, ← inner_conj_symm y, mul_conj, norm_sq_eq_def'], + push_cast, + ring +end + /-- -**Cauchy–Schwarz inequality**. This proof follows "Proof 2" on Wikipedia. +**Cauchy–Schwarz inequality**. We need this for the `core` structure to prove the triangle inequality below when showing the core is a normed group. -/ -lemma inner_mul_inner_self_le (x y : F) : abs ⟪x, y⟫ * abs ⟪y, x⟫ ≤ re ⟪x, x⟫ * re ⟪y, y⟫ := +lemma inner_mul_inner_self_le (x y : F) : ‖⟪x, y⟫‖ * ‖⟪y, x⟫‖ ≤ re ⟪x, x⟫ * re ⟪y, y⟫ := begin - by_cases hy : y = 0, - { rw [hy], simp only [is_R_or_C.abs_zero, inner_zero_left, mul_zero, add_monoid_hom.map_zero] }, - { change y ≠ 0 at hy, - have hy' : ⟪y, y⟫ ≠ 0 := inner_self_ne_zero.mpr hy, - set T := ⟪y, x⟫ / ⟪y, y⟫ with hT, - have h₁ : re ⟪y, x⟫ = re ⟪x, y⟫ := inner_re_symm _ _, - have h₂ : im ⟪y, x⟫ = -im ⟪x, y⟫ := inner_im_symm _ _, - have h₃ : ⟪y, x⟫ * ⟪x, y⟫ * ⟪y, y⟫ / (⟪y, y⟫ * ⟪y, y⟫) = ⟪y, x⟫ * ⟪x, y⟫ / ⟪y, y⟫, - { rw [mul_div_assoc], - have : ⟪y, y⟫ / (⟪y, y⟫ * ⟪y, y⟫) = 1 / ⟪y, y⟫ := - by rw [div_mul_eq_div_mul_one_div, div_self hy', one_mul], - rw [this, div_eq_mul_inv, one_mul, ←div_eq_mul_inv] }, - have h₄ : ⟪y, y⟫ = re ⟪y, y⟫ := by simp only [inner_self_re_to_K], - have h₅ : re ⟪y, y⟫ > 0, - { refine lt_of_le_of_ne inner_self_nonneg _, - intro H, - apply hy', - rw ext_iff, - exact ⟨by simp only [H, zero_re'], - by simp only [inner_self_im, add_monoid_hom.map_zero]⟩ }, - have h₆ : re ⟪y, y⟫ ≠ 0 := ne_of_gt h₅, - have hmain := calc - 0 ≤ re ⟪x - T • y, x - T • y⟫ - : inner_self_nonneg - ... = re ⟪x, x⟫ - re ⟪T • y, x⟫ - re ⟪x, T • y⟫ + re ⟪T • y, T • y⟫ - : by simp only [inner_sub_sub_self, inner_smul_left, inner_smul_right, h₁, h₂, - neg_mul, add_monoid_hom.map_add, mul_re, - conj_im, add_monoid_hom.map_sub, mul_neg, conj_re, neg_neg] - ... = re ⟪x, x⟫ - re (T† * ⟪y, x⟫) - re (T * ⟪x, y⟫) + re (T * T† * ⟪y, y⟫) - : by simp only [inner_smul_left, inner_smul_right, mul_assoc] - ... = re ⟪x, x⟫ - re (⟪x, y⟫ / ⟪y, y⟫ * ⟪y, x⟫) - : by field_simp [-mul_re, inner_conj_symm, hT, map_div₀, h₁, h₃] - ... = re ⟪x, x⟫ - re (⟪x, y⟫ * ⟪y, x⟫ / ⟪y, y⟫) - : by rw ←mul_div_right_comm - ... = re ⟪x, x⟫ - re (⟪x, y⟫ * ⟪y, x⟫ / re ⟪y, y⟫) - : by conv_lhs { rw [h₄] } - ... = re ⟪x, x⟫ - re (⟪x, y⟫ * ⟪y, x⟫) / re ⟪y, y⟫ - : by rw [div_re_of_real] - ... = re ⟪x, x⟫ - abs (⟪x, y⟫ * ⟪y, x⟫) / re ⟪y, y⟫ - : by rw [inner_mul_conj_re_abs] - ... = re ⟪x, x⟫ - abs ⟪x, y⟫ * abs ⟪y, x⟫ / re ⟪y, y⟫ - : by rw is_R_or_C.abs_mul, - have hmain' : abs ⟪x, y⟫ * abs ⟪y, x⟫ / re ⟪y, y⟫ ≤ re ⟪x, x⟫ := by linarith, - have := (mul_le_mul_right h₅).mpr hmain', - rwa [div_mul_cancel (abs ⟪x, y⟫ * abs ⟪y, x⟫) h₆] at this } + rcases eq_or_ne x 0 with (rfl | hx), + { simp only [inner_zero_left, map_zero, zero_mul, norm_zero] }, + { have hx' : 0 < norm_sqF x := inner_self_nonneg.lt_of_ne' (mt norm_sq_eq_zero.1 hx), + rw [← sub_nonneg, ← mul_nonneg_iff_right_nonneg_of_pos hx', ← norm_sq, ← norm_sq, + norm_inner_symm y, ← sq, ← cauchy_schwarz_aux], + exact inner_self_nonneg } end /-- Norm constructed from a `inner_product_space.core` structure, defined to be the square root @@ -304,18 +296,11 @@ by rw [norm_eq_sqrt_inner, ←sqrt_mul inner_self_nonneg (re ⟪x, x⟫), lemma sqrt_norm_sq_eq_norm (x : F) : sqrt (norm_sqF x) = ‖x‖ := rfl /-- Cauchy–Schwarz inequality with norm -/ -lemma abs_inner_le_norm (x y : F) : abs ⟪x, y⟫ ≤ ‖x‖ * ‖y‖ := -nonneg_le_nonneg_of_sq_le_sq (mul_nonneg (sqrt_nonneg _) (sqrt_nonneg _)) -begin - have H : ‖x‖ * ‖y‖ * (‖x‖ * ‖y‖) = re ⟪y, y⟫ * re ⟪x, x⟫, - { simp only [inner_self_eq_norm_mul_norm], ring, }, - rw H, - conv - begin - to_lhs, congr, rw [abs_inner_symm], - end, - exact inner_mul_inner_self_le y x, -end +lemma norm_inner_le_norm (x y : F) : ‖⟪x, y⟫‖ ≤ ‖x‖ * ‖y‖ := +nonneg_le_nonneg_of_sq_le_sq (mul_nonneg (sqrt_nonneg _) (sqrt_nonneg _)) $ + calc ‖⟪x, y⟫‖ * ‖⟪x, y⟫‖ = ‖⟪x, y⟫‖ * ‖⟪y, x⟫‖ : by rw [norm_inner_symm] + ... ≤ re ⟪x, x⟫ * re ⟪y, y⟫ : inner_mul_inner_self_le x y + ... = ‖x‖ * ‖y‖ * (‖x‖ * ‖y‖) : by simp only [inner_self_eq_norm_mul_norm]; ring /-- Normed group structure constructed from an `inner_product_space.core` structure -/ def to_normed_add_comm_group : normed_add_comm_group F := @@ -324,9 +309,9 @@ add_group_norm.to_normed_add_comm_group map_zero' := by simp only [sqrt_zero, inner_zero_right, map_zero], neg' := λ x, by simp only [inner_neg_left, neg_neg, inner_neg_right], add_le' := λ x y, begin - have h₁ : abs ⟪x, y⟫ ≤ ‖x‖ * ‖y‖ := abs_inner_le_norm _ _, - have h₂ : re ⟪x, y⟫ ≤ abs ⟪x, y⟫ := re_le_abs _, - have h₃ : re ⟪x, y⟫ ≤ ‖x‖ * ‖y‖ := by linarith, + have h₁ : ‖⟪x, y⟫‖ ≤ ‖x‖ * ‖y‖ := norm_inner_le_norm _ _, + have h₂ : re ⟪x, y⟫ ≤ ‖⟪x, y⟫‖ := (re_le_abs _).trans_eq (is_R_or_C.norm_eq_abs _).symm, + have h₃ : re ⟪x, y⟫ ≤ ‖x‖ * ‖y‖ := h₂.trans h₁, have h₄ : re ⟪y, x⟫ ≤ ‖x‖ * ‖y‖ := by rwa [←inner_conj_symm, conj_re], have : ‖x + y‖ * ‖x + y‖ ≤ (‖x‖ + ‖y‖) * (‖x‖ + ‖y‖), { simp only [←inner_self_eq_norm_mul_norm, inner_add_add_self, mul_add, mul_comm, map_add], @@ -342,16 +327,16 @@ def to_normed_space : normed_space 𝕜 F := { norm_smul_le := assume r x, begin rw [norm_eq_sqrt_inner, inner_smul_left, inner_smul_right, ←mul_assoc], - rw [conj_mul_eq_norm_sq_left, of_real_mul_re, sqrt_mul, ←inner_norm_sq_eq_inner_self, - of_real_re], + rw [is_R_or_C.conj_mul_eq_norm_sq_left, of_real_mul_re, sqrt_mul, ← coe_norm_sq_eq_inner_self, + of_real_re], { simp [sqrt_norm_sq_eq_norm, is_R_or_C.sqrt_norm_sq_eq_norm] }, { exact norm_sq_nonneg r } end } -end inner_product_space.of_core +end inner_product_space.core section -local attribute [instance] inner_product_space.of_core.to_normed_add_comm_group +local attribute [instance] inner_product_space.core.to_normed_add_comm_group /-- Given a `inner_product_space.core` structure on a space, one can use it to turn the space into an inner product space. The `normed_add_comm_group` structure is expected @@ -359,11 +344,11 @@ to already be defined with `inner_product_space.of_core.to_normed_add_comm_group def inner_product_space.of_core [add_comm_group F] [module 𝕜 F] (c : inner_product_space.core 𝕜 F) : inner_product_space 𝕜 F := begin - letI : normed_space 𝕜 F := @inner_product_space.of_core.to_normed_space 𝕜 F _ _ _ c, + letI : normed_space 𝕜 F := @inner_product_space.core.to_normed_space 𝕜 F _ _ _ c, exact { norm_sq_eq_inner := λ x, begin have h₁ : ‖x‖^2 = (sqrt (re (c.inner x x))) ^ 2 := rfl, - have h₂ : 0 ≤ re (c.inner x x) := inner_product_space.of_core.inner_self_nonneg, + have h₂ : 0 ≤ re (c.inner x x) := inner_product_space.core.inner_self_nonneg, simp [h₁, sq_sqrt, h₂], end, ..c } @@ -390,7 +375,7 @@ section basic_properties lemma real_inner_comm (x y : F) : ⟪y, x⟫_ℝ = ⟪x, y⟫_ℝ := @inner_conj_symm ℝ _ _ _ _ x y lemma inner_eq_zero_symm {x y : E} : ⟪x, y⟫ = 0 ↔ ⟪y, x⟫ = 0 := -⟨λ h, by simp [←inner_conj_symm, h], λ h, by simp [←inner_conj_symm, h]⟩ +by { rw [← inner_conj_symm], exact star_eq_zero } @[simp] lemma inner_self_im (x : E) : im ⟪x, x⟫ = 0 := by rw [← @of_real_inj 𝕜, im_eq_conj_sub]; simp @@ -489,7 +474,8 @@ lemma inner_re_zero_right (x : E) : re ⟪x, 0⟫ = 0 := by simp only [inner_zero_right, add_monoid_hom.map_zero] lemma inner_self_nonneg {x : E} : 0 ≤ re ⟪x, x⟫ := -by rw [←norm_sq_eq_inner]; exact pow_nonneg (norm_nonneg x) 2 +inner_product_space.to_core.nonneg_re x + lemma real_inner_self_nonneg {x : F} : 0 ≤ ⟪x, x⟫_ℝ := @inner_self_nonneg ℝ F _ _ _ x @[simp] lemma inner_self_re_to_K (x : E) : (re ⟪x, x⟫ : 𝕜) = ⟪x, x⟫ := @@ -587,60 +573,17 @@ lemma parallelogram_law {x y : E} : ⟪x + y, x + y⟫ + ⟪x - y, x - y⟫ = 2 * (⟪x, x⟫ + ⟪y, y⟫) := by simp [inner_add_add_self, inner_sub_sub_self, two_mul, sub_eq_add_neg, add_comm, add_left_comm] -/-- Cauchy–Schwarz inequality. This proof follows "Proof 2" on Wikipedia. -/ -lemma inner_mul_inner_self_le (x y : E) : abs ⟪x, y⟫ * abs ⟪y, x⟫ ≤ re ⟪x, x⟫ * re ⟪y, y⟫ := +/-- **Cauchy–Schwarz inequality**. -/ +lemma inner_mul_inner_self_le (x y : E) : ‖⟪x, y⟫‖ * ‖⟪y, x⟫‖ ≤ re ⟪x, x⟫ * re ⟪y, y⟫ := begin - by_cases hy : y = 0, - { rw [hy], simp only [is_R_or_C.abs_zero, inner_zero_left, mul_zero, add_monoid_hom.map_zero] }, - { have hy' : ⟪y, y⟫ ≠ 0 := inner_self_eq_zero.not.2 hy, - set T := ⟪y, x⟫ / ⟪y, y⟫ with hT, - have h₁ : re ⟪y, x⟫ = re ⟪x, y⟫ := inner_re_symm _ _, - have h₂ : im ⟪y, x⟫ = -im ⟪x, y⟫ := inner_im_symm _ _, - have h₃ : ⟪y, x⟫ * ⟪x, y⟫ * ⟪y, y⟫ / (⟪y, y⟫ * ⟪y, y⟫) = ⟪y, x⟫ * ⟪x, y⟫ / ⟪y, y⟫, - { rw [mul_div_assoc], - have : ⟪y, y⟫ / (⟪y, y⟫ * ⟪y, y⟫) = 1 / ⟪y, y⟫ := - by rw [div_mul_eq_div_mul_one_div, div_self hy', one_mul], - rw [this, div_eq_mul_inv, one_mul, ←div_eq_mul_inv] }, - have h₄ : ⟪y, y⟫ = re ⟪y, y⟫ := (inner_self_re_to_K _).symm, - have h₅ : re ⟪y, y⟫ > 0, - { refine lt_of_le_of_ne inner_self_nonneg _, - intro H, - apply hy', - rw is_R_or_C.ext_iff, - exact ⟨by simp only [H, zero_re'], - by simp only [inner_self_im, add_monoid_hom.map_zero]⟩ }, - have h₆ : re ⟪y, y⟫ ≠ 0 := ne_of_gt h₅, - have hmain := calc - 0 ≤ re ⟪x - T • y, x - T • y⟫ - : inner_self_nonneg - ... = re ⟪x, x⟫ - re ⟪T • y, x⟫ - re ⟪x, T • y⟫ + re ⟪T • y, T • y⟫ - : by simp only [inner_sub_sub_self, inner_smul_left, inner_smul_right, h₁, h₂, - neg_mul, add_monoid_hom.map_add, conj_im, - add_monoid_hom.map_sub, mul_neg, conj_re, neg_neg, mul_re] - ... = re ⟪x, x⟫ - re (T† * ⟪y, x⟫) - re (T * ⟪x, y⟫) + re (T * T† * ⟪y, y⟫) - : by simp only [inner_smul_left, inner_smul_right, mul_assoc] - ... = re ⟪x, x⟫ - re (⟪x, y⟫ / ⟪y, y⟫ * ⟪y, x⟫) - : by simp only [map_div₀, h₃, inner_conj_symm, sub_add_cancel] - with field_simps {discharger := tactic.field_simp.ne_zero} - ... = re ⟪x, x⟫ - re (⟪x, y⟫ * ⟪y, x⟫ / ⟪y, y⟫) - : by rw ←mul_div_right_comm - ... = re ⟪x, x⟫ - re (⟪x, y⟫ * ⟪y, x⟫ / re ⟪y, y⟫) - : by conv_lhs { rw [h₄] } - ... = re ⟪x, x⟫ - re (⟪x, y⟫ * ⟪y, x⟫) / re ⟪y, y⟫ - : by rw [div_re_of_real] - ... = re ⟪x, x⟫ - abs (⟪x, y⟫ * ⟪y, x⟫) / re ⟪y, y⟫ - : by rw [inner_mul_conj_re_abs] - ... = re ⟪x, x⟫ - abs ⟪x, y⟫ * abs ⟪y, x⟫ / re ⟪y, y⟫ - : by rw is_R_or_C.abs_mul, - have hmain' : abs ⟪x, y⟫ * abs ⟪y, x⟫ / re ⟪y, y⟫ ≤ re ⟪x, x⟫ := by linarith, - have := (mul_le_mul_right h₅).mpr hmain', - rwa [div_mul_cancel (abs ⟪x, y⟫ * abs ⟪y, x⟫) h₆] at this } + letI c : inner_product_space.core 𝕜 E := inner_product_space.to_core, + exact inner_product_space.core.inner_mul_inner_self_le x y end /-- Cauchy–Schwarz inequality for real inner products. -/ lemma real_inner_mul_inner_self_le (x y : F) : ⟪x, y⟫_ℝ * ⟪x, y⟫_ℝ ≤ ⟪x, x⟫_ℝ * ⟪y, y⟫_ℝ := -calc ⟪x, y⟫_ℝ * ⟪x, y⟫_ℝ ≤ is_R_or_C.abs ⟪x, y⟫_ℝ * is_R_or_C.abs ⟪y, x⟫_ℝ : - by { rw [real_inner_comm y, ← is_R_or_C.abs_mul, ← is_R_or_C.norm_eq_abs], exact le_abs_self _ } +calc ⟪x, y⟫_ℝ * ⟪x, y⟫_ℝ ≤ ‖⟪x, y⟫_ℝ‖ * ‖⟪y, x⟫_ℝ‖ : + by { rw [real_inner_comm y, ← norm_mul], exact le_abs_self _ } ... ≤ ⟪x, x⟫_ℝ * ⟪y, y⟫_ℝ : @inner_mul_inner_self_le ℝ _ _ _ _ x y /-- A family of vectors is linearly independent if they are nonzero @@ -927,7 +870,7 @@ calc ‖x‖ = sqrt (‖x‖ ^ 2) : (sqrt_sq (norm_nonneg _)).symm ... = sqrt (re ⟪x, x⟫) : congr_arg _ (norm_sq_eq_inner _) lemma norm_eq_sqrt_real_inner (x : F) : ‖x‖ = sqrt ⟪x, x⟫_ℝ := -by { have h := @norm_eq_sqrt_inner ℝ F _ _ _ x, simpa using h } +@norm_eq_sqrt_inner ℝ _ _ _ _ x lemma inner_self_eq_norm_mul_norm (x : E) : re ⟪x, x⟫ = ‖x‖ * ‖x‖ := by rw [@norm_eq_sqrt_inner 𝕜, ←sqrt_mul inner_self_nonneg (re ⟪x, x⟫), @@ -990,18 +933,15 @@ lemma norm_sub_mul_self_real (x y : F) : ‖x - y‖ * ‖x - y‖ = ‖x‖ * by { have h := @norm_sub_mul_self ℝ _ _ _ _ x y, simpa using h } /-- Cauchy–Schwarz inequality with norm -/ -lemma abs_inner_le_norm (x y : E) : abs ⟪x, y⟫ ≤ ‖x‖ * ‖y‖ := -nonneg_le_nonneg_of_sq_le_sq (mul_nonneg (norm_nonneg _) (norm_nonneg _)) +lemma norm_inner_le_norm (x y : E) : ‖⟪x, y⟫‖ ≤ ‖x‖ * ‖y‖ := begin - have : ‖x‖ * ‖y‖ * (‖x‖ * ‖y‖) = (re ⟪x, x⟫) * (re ⟪y, y⟫), - simp only [inner_self_eq_norm_mul_norm], ring, - rw this, - conv_lhs { congr, skip, rw [abs_inner_symm] }, - exact inner_mul_inner_self_le _ _ + rw [norm_eq_sqrt_inner x, norm_eq_sqrt_inner y], + letI : inner_product_space.core 𝕜 E := inner_product_space.to_core, + exact inner_product_space.core.norm_inner_le_norm x y end -lemma norm_inner_le_norm (x y : E) : ‖⟪x, y⟫‖ ≤ ‖x‖ * ‖y‖ := -(is_R_or_C.norm_eq_abs _).le.trans (abs_inner_le_norm x y) +lemma abs_inner_le_norm (x y : E) : abs ⟪x, y⟫ ≤ ‖x‖ * ‖y‖ := +(norm_eq_abs ⟪x, y⟫).symm.trans_le (norm_inner_le_norm _ _) lemma nnnorm_inner_le_nnnorm (x y : E) : ‖⟪x, y⟫‖₊ ≤ ‖x‖₊ * ‖y‖₊ := norm_inner_le_norm x y @@ -1010,8 +950,8 @@ lemma re_inner_le_norm (x y : E) : re ⟪x, y⟫ ≤ ‖x‖ * ‖y‖ := le_trans (re_le_abs (inner x y)) (abs_inner_le_norm x y) /-- Cauchy–Schwarz inequality with norm -/ -lemma abs_real_inner_le_norm (x y : F) : absR ⟪x, y⟫_ℝ ≤ ‖x‖ * ‖y‖ := -by { have h := @abs_inner_le_norm ℝ F _ _ _ x y, simpa using h } +lemma abs_real_inner_le_norm (x y : F) : |⟪x, y⟫_ℝ| ≤ ‖x‖ * ‖y‖ := +(real.norm_eq_abs _).ge.trans (norm_inner_le_norm x y) /-- Cauchy–Schwarz inequality with norm -/ lemma real_inner_le_norm (x y : F) : ⟪x, y⟫_ℝ ≤ ‖x‖ * ‖y‖ := @@ -1437,64 +1377,63 @@ itself, divided by the product of their norms, has value -1. -/ lemma real_inner_div_norm_mul_norm_eq_neg_one_of_ne_zero_of_neg_mul {x : F} {r : ℝ} (hx : x ≠ 0) (hr : r < 0) : ⟪x, r • x⟫_ℝ / (‖x‖ * ‖r • x‖) = -1 := begin - rw [real_inner_smul_self_right, norm_smul, real.norm_eq_abs, ←mul_assoc ‖x‖, mul_comm _ (absR r), + rw [real_inner_smul_self_right, norm_smul, real.norm_eq_abs, ←mul_assoc ‖x‖, mul_comm _ (|r|), mul_assoc, abs_of_neg hr, neg_mul, div_neg_eq_neg_div, div_self], - exact mul_ne_zero (ne_of_lt hr) - (λ h, hx (norm_eq_zero.1 (eq_zero_of_mul_self_eq_zero h))) + exact mul_ne_zero hr.ne (mul_self_ne_zero.2 (norm_ne_zero_iff.2 hx)) +end + +lemma norm_inner_eq_norm_tfae (x y : E) : + tfae [‖⟪x, y⟫‖ = ‖x‖ * ‖y‖, + x = 0 ∨ y = (⟪x, y⟫ / ⟪x, x⟫) • x, + x = 0 ∨ ∃ r : 𝕜, y = r • x, + x = 0 ∨ y ∈ 𝕜 ∙ x] := +begin + tfae_have : 1 → 2, + { refine λ h, or_iff_not_imp_left.2 (λ hx₀, _), + have : ‖x‖ ^ 2 ≠ 0 := pow_ne_zero _ (norm_ne_zero_iff.2 hx₀), + rw [← sq_eq_sq (norm_nonneg _) (mul_nonneg (norm_nonneg _) (norm_nonneg _)), + mul_pow, ← mul_right_inj' this, eq_comm, ← sub_eq_zero, ← mul_sub] at h, + simp only [@norm_sq_eq_inner 𝕜] at h, + letI : inner_product_space.core 𝕜 E := inner_product_space.to_core, + erw [← inner_product_space.core.cauchy_schwarz_aux, + inner_product_space.core.norm_sq_eq_zero, sub_eq_zero] at h, + rw [div_eq_inv_mul, mul_smul, h, inv_smul_smul₀], + rwa [inner_self_ne_zero] }, + tfae_have : 2 → 3, from λ h, h.imp_right (λ h', ⟨_, h'⟩), + tfae_have : 3 → 1, + { rintro (rfl | ⟨r, rfl⟩); simp [inner_smul_right, norm_smul, inner_self_eq_norm_sq_to_K, + inner_self_eq_norm_mul_norm, sq, mul_left_comm] }, + tfae_have : 3 ↔ 4, by simp only [submodule.mem_span_singleton, eq_comm], + tfae_finish end +/-- +If the inner product of two vectors is equal to the product of their norms, then the two vectors +are multiples of each other. One form of the equality case for Cauchy-Schwarz. +Compare `inner_eq_norm_mul_iff`, which takes the stronger hypothesis `⟪x, y⟫ = ‖x‖ * ‖y‖`. -/ +lemma norm_inner_eq_norm_iff {x y : E} (hx₀ : x ≠ 0) (hy₀ : y ≠ 0) : + ‖⟪x, y⟫‖ = ‖x‖ * ‖y‖ ↔ ∃ (r : 𝕜), r ≠ 0 ∧ y = r • x := +calc ‖⟪x, y⟫‖ = ‖x‖ * ‖y‖ ↔ x = 0 ∨ ∃ r : 𝕜, y = r • x : + (@norm_inner_eq_norm_tfae 𝕜 _ _ _ _ x y).out 0 2 +... ↔ ∃ r : 𝕜, y = r • x : or_iff_right hx₀ +... ↔ ∃ r : 𝕜, r ≠ 0 ∧ y = r • x : + ⟨λ ⟨r, h⟩, ⟨r, λ hr₀, hy₀ $ h.symm ▸ smul_eq_zero.2 $ or.inl hr₀, h⟩, λ ⟨r, hr₀, h⟩, ⟨r, h⟩⟩ + /-- The inner product of two vectors, divided by the product of their norms, has absolute value 1 if and only if they are nonzero and one is a multiple of the other. One form of equality case for Cauchy-Schwarz. -/ -lemma abs_inner_div_norm_mul_norm_eq_one_iff (x y : E) : - abs (⟪x, y⟫ / (‖x‖ * ‖y‖)) = 1 ↔ (x ≠ 0 ∧ ∃ (r : 𝕜), r ≠ 0 ∧ y = r • x) := +lemma norm_inner_div_norm_mul_norm_eq_one_iff (x y : E) : + ‖(⟪x, y⟫ / (‖x‖ * ‖y‖))‖ = 1 ↔ (x ≠ 0 ∧ ∃ (r : 𝕜), r ≠ 0 ∧ y = r • x) := begin split, { intro h, - have hx0 : x ≠ 0, - { intro hx0, - rw [hx0, inner_zero_left, zero_div] at h, - norm_num at h, }, - refine and.intro hx0 _, - set r := ⟪x, y⟫ / (‖x‖ * ‖x‖) with hr, - use r, - set t := y - r • x with ht, - have ht0 : ⟪x, t⟫ = 0, - { rw [ht, inner_sub_right, inner_smul_right, hr], - norm_cast, - rw [←@inner_self_eq_norm_mul_norm 𝕜, inner_self_re_to_K, - div_mul_cancel _ (λ h, hx0 ((@inner_self_eq_zero 𝕜 _ _ _ _ _).1 h)), sub_self] }, - replace h : ‖r • x‖ / ‖t + r • x‖ = 1, - { rw [←sub_add_cancel y (r • x), ←ht, inner_add_right, ht0, zero_add, inner_smul_right, - is_R_or_C.abs_div, is_R_or_C.abs_mul, ←inner_self_re_abs, - inner_self_eq_norm_mul_norm] at h, - norm_cast at h, - rwa [_root_.abs_mul, abs_norm, abs_norm, ←mul_assoc, mul_comm, - mul_div_mul_left _ _ (λ h, hx0 (norm_eq_zero.1 h)), ←is_R_or_C.norm_eq_abs, - ←norm_smul] at h }, - have hr0 : r ≠ 0, - { intro hr0, - rw [hr0, zero_smul, norm_zero, zero_div] at h, - norm_num at h }, - refine and.intro hr0 _, - have h2 : ‖r • x‖ ^ 2 = ‖t + r • x‖ ^ 2, - { rw [eq_of_div_eq_one h] }, - replace h2 : ⟪r • x, r • x⟫ = ⟪t, t⟫ + ⟪t, r • x⟫ + ⟪r • x, t⟫ + ⟪r • x, r • x⟫, - { rw [sq, sq, ←@inner_self_eq_norm_mul_norm 𝕜, ←@inner_self_eq_norm_mul_norm 𝕜] at h2, - have h2' := congr_arg (λ z : ℝ, (z : 𝕜)) h2, - simp_rw [inner_self_re_to_K, inner_add_add_self] at h2', - exact h2' }, - conv at h2 in ⟪r • x, t⟫ { rw [inner_smul_left, ht0, mul_zero] }, - symmetry' at h2, - have h₁ : ⟪t, r • x⟫ = 0 := by { rw [inner_smul_right, ←inner_conj_symm, ht0], simp }, - rw [add_zero, h₁, add_left_eq_self, add_zero, inner_self_eq_zero] at h2, - rw h2 at ht, - exact eq_of_sub_eq_zero ht.symm }, - { intro h, - rcases h with ⟨hx, ⟨r, ⟨hr, hy⟩⟩⟩, - rw [hy, is_R_or_C.abs_div], - norm_cast, - rw [_root_.abs_mul, abs_norm, abs_norm], + have hx₀ : x ≠ 0 := λ h₀, by simpa [h₀] using h, + have hy₀ : y ≠ 0 := λ h₀, by simpa [h₀] using h, + refine ⟨hx₀, (norm_inner_eq_norm_iff hx₀ hy₀).1 $ eq_of_div_eq_one _⟩, + simpa using h }, + { rintro ⟨hx, ⟨r, ⟨hr, rfl⟩⟩⟩, + simp only [is_R_or_C.abs_div, is_R_or_C.abs_mul, is_R_or_C.abs_of_real, is_R_or_C.norm_eq_abs, + abs_norm], exact abs_inner_div_norm_mul_norm_eq_one_of_ne_zero_of_ne_zero_mul hx hr } end @@ -1502,27 +1441,43 @@ end norms, has absolute value 1 if and only if they are nonzero and one is a multiple of the other. One form of equality case for Cauchy-Schwarz. -/ lemma abs_real_inner_div_norm_mul_norm_eq_one_iff (x y : F) : - absR (⟪x, y⟫_ℝ / (‖x‖ * ‖y‖)) = 1 ↔ (x ≠ 0 ∧ ∃ (r : ℝ), r ≠ 0 ∧ y = r • x) := + |⟪x, y⟫_ℝ / (‖x‖ * ‖y‖)| = 1 ↔ (x ≠ 0 ∧ ∃ (r : ℝ), r ≠ 0 ∧ y = r • x) := +@norm_inner_div_norm_mul_norm_eq_one_iff ℝ F _ _ _ x y + +lemma inner_eq_norm_mul_iff_div {x y : E} (h₀ : x ≠ 0) : + ⟪x, y⟫ = (‖x‖ : 𝕜) * ‖y‖ ↔ (‖y‖ / ‖x‖ : 𝕜) • x = y := begin - have := @abs_inner_div_norm_mul_norm_eq_one_iff ℝ F _ _ _ x y, - simpa [coe_real_eq_id] using this, + have h₀' := h₀, + rw [← norm_ne_zero_iff, ne.def, ← @of_real_eq_zero 𝕜] at h₀', + split; intro h, + { have : x = 0 ∨ y = (⟪x, y⟫ / ⟪x, x⟫ : 𝕜) • x := + ((@norm_inner_eq_norm_tfae 𝕜 _ _ _ _ x y).out 0 1).1 (by simp [h]), + rw [this.resolve_left h₀, h], + simp [norm_smul, is_R_or_C.norm_eq_abs, inner_self_abs_to_K, h₀'] }, + { conv_lhs { rw [← h, inner_smul_right, inner_self_eq_norm_sq_to_K] }, + field_simp [sq, mul_left_comm] } end -/-- -If the inner product of two vectors is equal to the product of their norms, then the two vectors -are multiples of each other. One form of the equality case for Cauchy-Schwarz. -Compare `inner_eq_norm_mul_iff`, which takes the stronger hypothesis `⟪x, y⟫ = ‖x‖ * ‖y‖`. -/ -lemma abs_inner_eq_norm_iff (x y : E) (hx0 : x ≠ 0) (hy0 : y ≠ 0): - abs ⟪x, y⟫ = ‖x‖ * ‖y‖ ↔ ∃ (r : 𝕜), r ≠ 0 ∧ y = r • x := +/-- If the inner product of two vectors is equal to the product of their norms (i.e., +`⟪x, y⟫ = ‖x‖ * ‖y‖`), then the two vectors are nonnegative real multiples of each other. One form +of the equality case for Cauchy-Schwarz. +Compare `norm_inner_eq_norm_iff`, which takes the weaker hypothesis `abs ⟪x, y⟫ = ‖x‖ * ‖y‖`. -/ +lemma inner_eq_norm_mul_iff {x y : E} : + ⟪x, y⟫ = (‖x‖ : 𝕜) * ‖y‖ ↔ (‖y‖ : 𝕜) • x = (‖x‖ : 𝕜) • y := begin - have hxy0 : ‖x‖ * ‖y‖ ≠ 0 := mul_ne_zero (norm_eq_zero.not.2 hx0) (norm_eq_zero.not.2 hy0), - have h₁ : abs ⟪x, y⟫ = ‖x‖ * ‖y‖ ↔ abs (⟪x, y⟫ / (‖x‖ * ‖y‖)) = 1, - { rw [←algebra_map.coe_mul, is_R_or_C.abs_div, is_R_or_C.abs_of_nonneg, div_eq_one_iff_eq hxy0], - positivity }, - rw [h₁, abs_inner_div_norm_mul_norm_eq_one_iff x y], - exact and_iff_right hx0, + rcases eq_or_ne x 0 with (rfl | h₀), + { simp }, + { rw [inner_eq_norm_mul_iff_div h₀, div_eq_inv_mul, mul_smul, inv_smul_eq_iff₀], + rwa [ne.def, of_real_eq_zero, norm_eq_zero] }, end +/-- If the inner product of two vectors is equal to the product of their norms (i.e., +`⟪x, y⟫ = ‖x‖ * ‖y‖`), then the two vectors are nonnegative real multiples of each other. One form +of the equality case for Cauchy-Schwarz. +Compare `norm_inner_eq_norm_iff`, which takes the weaker hypothesis `abs ⟪x, y⟫ = ‖x‖ * ‖y‖`. -/ +lemma inner_eq_norm_mul_iff_real {x y : F} : ⟪x, y⟫_ℝ = ‖x‖ * ‖y‖ ↔ ‖y‖ • x = ‖x‖ • y := +inner_eq_norm_mul_iff + /-- The inner product of two vectors, divided by the product of their norms, has value 1 if and only if they are nonzero and one is a positive multiple of the other. -/ @@ -1531,20 +1486,11 @@ lemma real_inner_div_norm_mul_norm_eq_one_iff (x y : F) : begin split, { intro h, - have ha := h, - apply_fun absR at ha, - norm_num at ha, - rcases (abs_real_inner_div_norm_mul_norm_eq_one_iff x y).1 ha with ⟨hx, ⟨r, ⟨hr, hy⟩⟩⟩, - use [hx, r], - refine and.intro _ hy, - by_contradiction hrneg, - rw hy at h, - rw real_inner_div_norm_mul_norm_eq_neg_one_of_ne_zero_of_neg_mul hx - (lt_of_le_of_ne (le_of_not_lt hrneg) hr) at h, - norm_num at h }, - { intro h, - rcases h with ⟨hx, ⟨r, ⟨hr, hy⟩⟩⟩, - rw hy, + have hx₀ : x ≠ 0 := λ h₀, by simpa [h₀] using h, + have hy₀ : y ≠ 0 := λ h₀, by simpa [h₀] using h, + refine ⟨hx₀, ‖y‖ / ‖x‖, div_pos (norm_pos_iff.2 hy₀) (norm_pos_iff.2 hx₀), _⟩, + exact ((inner_eq_norm_mul_iff_div hx₀).1 (eq_of_div_eq_one h)).symm }, + { rintro ⟨hx, ⟨r, ⟨hr, rfl⟩⟩⟩, exact real_inner_div_norm_mul_norm_eq_one_of_ne_zero_of_pos_mul hx hr } end @@ -1560,48 +1506,9 @@ begin rw [neg_pos, neg_smul, neg_inj] end -/-- If the inner product of two vectors is equal to the product of their norms (i.e., -`⟪x, y⟫ = ‖x‖ * ‖y‖`), then the two vectors are nonnegative real multiples of each other. One form -of the equality case for Cauchy-Schwarz. -Compare `abs_inner_eq_norm_iff`, which takes the weaker hypothesis `abs ⟪x, y⟫ = ‖x‖ * ‖y‖`. -/ -lemma inner_eq_norm_mul_iff {x y : E} : - ⟪x, y⟫ = (‖x‖ : 𝕜) * ‖y‖ ↔ (‖y‖ : 𝕜) • x = (‖x‖ : 𝕜) • y := -begin - by_cases h : (x = 0 ∨ y = 0), -- WLOG `x` and `y` are nonzero - { cases h; simp [h] }, - calc ⟪x, y⟫ = (‖x‖ : 𝕜) * ‖y‖ ↔ ‖x‖ * ‖y‖ = re ⟪x, y⟫ : - begin - norm_cast, - split, - { intros h', - simp [h'] }, - { have cauchy_schwarz := abs_inner_le_norm x y, - intros h', - rw h' at ⊢ cauchy_schwarz, - rwa re_eq_self_of_le } - end - ... ↔ 2 * ‖x‖ * ‖y‖ * (‖x‖ * ‖y‖ - re ⟪x, y⟫) = 0 : - by simp [h, show (2:ℝ) ≠ 0, by norm_num, sub_eq_zero] - ... ↔ ‖(‖y‖:𝕜) • x - (‖x‖:𝕜) • y‖ * ‖(‖y‖:𝕜) • x - (‖x‖:𝕜) • y‖ = 0 : - begin - simp only [@norm_sub_mul_self 𝕜, inner_smul_left, inner_smul_right, norm_smul, conj_of_real, - is_R_or_C.norm_eq_abs, abs_of_real, of_real_im, of_real_re, mul_re, abs_norm], - refine eq.congr _ rfl, - ring - end - ... ↔ (‖y‖ : 𝕜) • x = (‖x‖ : 𝕜) • y : by simp [norm_sub_eq_zero_iff] -end - -/-- If the inner product of two vectors is equal to the product of their norms (i.e., -`⟪x, y⟫ = ‖x‖ * ‖y‖`), then the two vectors are nonnegative real multiples of each other. One form -of the equality case for Cauchy-Schwarz. -Compare `abs_inner_eq_norm_iff`, which takes the weaker hypothesis `abs ⟪x, y⟫ = ‖x‖ * ‖y‖`. -/ -lemma inner_eq_norm_mul_iff_real {x y : F} : ⟪x, y⟫_ℝ = ‖x‖ * ‖y‖ ↔ ‖y‖ • x = ‖x‖ • y := -inner_eq_norm_mul_iff - /-- If the inner product of two unit vectors is `1`, then the two vectors are equal. One form of the equality case for Cauchy-Schwarz. -/ -lemma inner_eq_norm_mul_iff_of_norm_one {x y : E} (hx : ‖x‖ = 1) (hy : ‖y‖ = 1) : +lemma inner_eq_one_iff_of_norm_one {x y : E} (hx : ‖x‖ = 1) (hy : ‖y‖ = 1) : ⟪x, y⟫ = 1 ↔ x = y := by { convert inner_eq_norm_mul_iff using 2; simp [hx, hy] } diff --git a/src/analysis/quaternion.lean b/src/analysis/quaternion.lean index 35bf8ff9e3bd3..f00c116fe91f6 100644 --- a/src/analysis/quaternion.lean +++ b/src/analysis/quaternion.lean @@ -42,8 +42,8 @@ lemma inner_self (a : ℍ) : ⟪a, a⟫ = norm_sq a := rfl lemma inner_def (a b : ℍ) : ⟪a, b⟫ = (a * star b).re := rfl noncomputable instance : normed_add_comm_group ℍ := -@inner_product_space.of_core.to_normed_add_comm_group ℝ ℍ _ _ _ -{ inner := has_inner.inner, +@inner_product_space.core.to_normed_add_comm_group ℝ ℍ _ _ _ +{ to_has_inner := infer_instance, conj_symm := λ x y, by simp [inner_def, mul_comm], nonneg_re := λ x, norm_sq_nonneg, definite := λ x, norm_sq_eq_zero.1, diff --git a/src/geometry/manifold/instances/sphere.lean b/src/geometry/manifold/instances/sphere.lean index c8ad586833bae..6ab30af5703da 100644 --- a/src/geometry/manifold/instances/sphere.lean +++ b/src/geometry/manifold/instances/sphere.lean @@ -283,7 +283,7 @@ def stereographic (hv : ‖v‖ = 1) : local_homeomorph (sphere (0:E) 1) (ℝ open_source := is_open_compl_singleton, open_target := is_open_univ, continuous_to_fun := continuous_on_stereo_to_fun.comp continuous_subtype_coe.continuous_on - (λ w h, h ∘ subtype.ext ∘ eq.symm ∘ (inner_eq_norm_mul_iff_of_norm_one hv (by simp)).mp), + (λ w h, h ∘ subtype.ext ∘ eq.symm ∘ (inner_eq_one_iff_of_norm_one hv (by simp)).mp), continuous_inv_fun := (continuous_stereo_inv_fun hv).continuous_on } lemma stereographic_apply (hv : ‖v‖ = 1) (x : sphere (0 : E) 1) : @@ -363,7 +363,7 @@ section smooth_manifold lemma sphere_ext_iff (u v : sphere (0:E) 1) : u = v ↔ ⟪(u:E), v⟫_ℝ = 1 := -by simp [subtype.ext_iff, inner_eq_norm_mul_iff_of_norm_one] +by simp [subtype.ext_iff, inner_eq_one_iff_of_norm_one] lemma stereographic'_symm_apply {n : ℕ} [fact (finrank ℝ E = n + 1)] (v : sphere (0:E) 1) (x : euclidean_space ℝ (fin n)) : @@ -448,7 +448,7 @@ begin ext x, have hfxv : f x = -↑v ↔ ⟪f x, -↑v⟫_ℝ = 1, { have hfx : ‖f x‖ = 1 := by simpa using hf' x, - rw inner_eq_norm_mul_iff_of_norm_one hfx, + rw inner_eq_one_iff_of_norm_one hfx, exact norm_eq_of_mem_sphere (-v) }, dsimp [chart_at], simp [not_iff_not, subtype.ext_iff, hfxv, real_inner_comm] diff --git a/src/linear_algebra/matrix/pos_def.lean b/src/linear_algebra/matrix/pos_def.lean index c814a4996c235..6e224830d8571 100644 --- a/src/linear_algebra/matrix/pos_def.lean +++ b/src/linear_algebra/matrix/pos_def.lean @@ -143,9 +143,9 @@ variables {𝕜 : Type*} [is_R_or_C 𝕜] {n : Type*} [fintype n] @[reducible] noncomputable def normed_add_comm_group.of_matrix {M : matrix n n 𝕜} (hM : M.pos_def) : normed_add_comm_group (n → 𝕜) := -@inner_product_space.of_core.to_normed_add_comm_group _ _ _ _ _ +@inner_product_space.core.to_normed_add_comm_group _ _ _ _ _ { inner := λ x y, dot_product (star x) (M.mul_vec y), - conj_symm := λ x y, by + conj_symm := λ x y, by dsimp only [has_inner.inner]; rw [star_dot_product, star_ring_end_apply, star_star, star_mul_vec, dot_product_mul_vec, hM.is_hermitian.eq], nonneg_re := λ x, @@ -154,10 +154,10 @@ noncomputable def normed_add_comm_group.of_matrix {M : matrix n n 𝕜} (hM : M. { simp [h] }, { exact le_of_lt (hM.2 x h) } end, - definite := λ x hx, + definite := λ x (hx : dot_product _ _ = 0), begin by_contra' h, - simpa [hx, lt_self_iff_false] using hM.2 x h, + simpa [hx, lt_irrefl] using hM.2 x h, end, add_left := by simp only [star_add, add_dot_product, eq_self_iff_true, forall_const], smul_left := λ x y r, by rw [← smul_eq_mul, ←smul_dot_product, star_ring_end_apply, ← star_smul] } From 698c80e908b68f7cd37f07e0bf4669b41cfe13ae Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Thu, 4 May 2023 11:37:56 +0000 Subject: [PATCH 0949/1291] chore(probability/kernel/composition): swap the order of the arguments of prod_mk_left (#18929) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Since `prod_mk_left γ κ` creates a `kernel (γ × α) β` from `κ : kernel α β`, it makes more sense to put the `γ` argument to the left. --- src/probability/kernel/composition.lean | 20 ++++++++++---------- 1 file changed, 10 insertions(+), 10 deletions(-) diff --git a/src/probability/kernel/composition.lean b/src/probability/kernel/composition.lean index 75b1c1930c9cb..569398e57c89b 100644 --- a/src/probability/kernel/composition.lean +++ b/src/probability/kernel/composition.lean @@ -462,7 +462,7 @@ open_locale probability_theory section fst_snd /-- Define a `kernel (γ × α) β` from a `kernel α β` by taking the comap of the projection. -/ -def prod_mk_left (κ : kernel α β) (γ : Type*) [measurable_space γ] : kernel (γ × α) β := +def prod_mk_left (γ : Type*) [measurable_space γ] (κ : kernel α β) : kernel (γ × α) β := comap κ prod.snd measurable_snd variables {γ : Type*} {mγ : measurable_space γ} {f : β → γ} {g : γ → α} @@ -470,24 +470,24 @@ variables {γ : Type*} {mγ : measurable_space γ} {f : β → γ} {g : γ → include mγ lemma prod_mk_left_apply (κ : kernel α β) (ca : γ × α) : - prod_mk_left κ γ ca = κ ca.snd := rfl + prod_mk_left γ κ ca = κ ca.snd := rfl lemma prod_mk_left_apply' (κ : kernel α β) (ca : γ × α) (s : set β) : - prod_mk_left κ γ ca s = κ ca.snd s := rfl + prod_mk_left γ κ ca s = κ ca.snd s := rfl lemma lintegral_prod_mk_left (κ : kernel α β) (ca : γ × α) (g : β → ℝ≥0∞) : - ∫⁻ b, g b ∂(prod_mk_left κ γ ca) = ∫⁻ b, g b ∂(κ ca.snd) := rfl + ∫⁻ b, g b ∂(prod_mk_left γ κ ca) = ∫⁻ b, g b ∂(κ ca.snd) := rfl instance is_markov_kernel.prod_mk_left (κ : kernel α β) [is_markov_kernel κ] : - is_markov_kernel (prod_mk_left κ γ) := + is_markov_kernel (prod_mk_left γ κ) := by { rw prod_mk_left, apply_instance, } instance is_finite_kernel.prod_mk_left (κ : kernel α β) [is_finite_kernel κ] : - is_finite_kernel (prod_mk_left κ γ) := + is_finite_kernel (prod_mk_left γ κ) := by { rw prod_mk_left, apply_instance, } instance is_s_finite_kernel.prod_mk_left (κ : kernel α β) [is_s_finite_kernel κ] : - is_s_finite_kernel (prod_mk_left κ γ) := + is_s_finite_kernel (prod_mk_left γ κ) := by { rw prod_mk_left, apply_instance, } /-- Define a `kernel (β × α) γ` from a `kernel (α × β) γ` by taking the comap of `prod.swap`. -/ @@ -613,7 +613,7 @@ include mγ noncomputable def comp (η : kernel β γ) [is_s_finite_kernel η] (κ : kernel α β) [is_s_finite_kernel κ] : kernel α γ := -snd (κ ⊗ₖ prod_mk_left η α) +snd (κ ⊗ₖ prod_mk_left α η) localized "infix (name := kernel.comp) ` ∘ₖ `:100 := probability_theory.kernel.comp" in probability_theory @@ -632,7 +632,7 @@ lemma lintegral_comp (η : kernel β γ) [is_s_finite_kernel η] (κ : kernel α ∫⁻ c, g c ∂((η ∘ₖ κ) a) = ∫⁻ b, ∫⁻ c, g c ∂(η b) ∂(κ a) := begin rw [comp, lintegral_snd _ _ hg], - change ∫⁻ bc, (λ a b, g b) bc.fst bc.snd ∂((κ ⊗ₖ prod_mk_left η α) a) + change ∫⁻ bc, (λ a b, g b) bc.fst bc.snd ∂((κ ⊗ₖ prod_mk_left α η) a) = ∫⁻ b, ∫⁻ c, g c ∂(η b) ∂(κ a), exact lintegral_comp_prod _ _ _ (hg.comp measurable_snd), end @@ -691,7 +691,7 @@ include mγ noncomputable def prod (κ : kernel α β) [is_s_finite_kernel κ] (η : kernel α γ) [is_s_finite_kernel η] : kernel α (β × γ) := -κ ⊗ₖ (swap_left (prod_mk_left η β)) +κ ⊗ₖ (swap_left (prod_mk_left β η)) localized "infix (name := kernel.prod) ` ×ₖ `:100 := probability_theory.kernel.prod" in probability_theory From caf83ba4dfbf4e2f28e1ae6a0780cbafc3d19d6f Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Thu, 4 May 2023 11:37:57 +0000 Subject: [PATCH 0950/1291] chore(probability/kernel/basic): make the function argument of kernel.deterministic explicit (#18930) If that argument is implicit, the infoview often shows only `deterministic _`, which does not allow to see which function is used. Making it explicit fixes that problem. --- src/probability/kernel/basic.lean | 10 +++++----- src/probability/kernel/composition.lean | 4 ++-- 2 files changed, 7 insertions(+), 7 deletions(-) diff --git a/src/probability/kernel/basic.lean b/src/probability/kernel/basic.lean index 6ebb39aace833..a3163ffd2284f 100644 --- a/src/probability/kernel/basic.lean +++ b/src/probability/kernel/basic.lean @@ -28,7 +28,7 @@ Classes of kernels: * `is_s_finite_kernel κ`: a kernel is called s-finite if it is a countable sum of finite kernels. Particular kernels: -* `deterministic {f : α → β} (hf : measurable f)`: kernel `a ↦ measure.dirac (f a)`. +* `deterministic (f : α → β) (hf : measurable f)`: kernel `a ↦ measure.dirac (f a)`. * `const α (μβ : measure β)`: constant kernel `a ↦ μβ`. * `kernel.restrict κ (hs : measurable_set s)`: kernel for which the image of `a : α` is `(κ a).restrict s`. @@ -317,7 +317,7 @@ section deterministic /-- Kernel which to `a` associates the dirac measure at `f a`. This is a Markov kernel. -/ noncomputable -def deterministic {f : α → β} (hf : measurable f) : +def deterministic (f : α → β) (hf : measurable f) : kernel α β := { val := λ a, measure.dirac (f a), property := @@ -328,11 +328,11 @@ def deterministic {f : α → β} (hf : measurable f) : end, } lemma deterministic_apply {f : α → β} (hf : measurable f) (a : α) : - deterministic hf a = measure.dirac (f a) := rfl + deterministic f hf a = measure.dirac (f a) := rfl lemma deterministic_apply' {f : α → β} (hf : measurable f) (a : α) {s : set β} (hs : measurable_set s) : - deterministic hf a s = s.indicator (λ _, 1) (f a) := + deterministic f hf a s = s.indicator (λ _, 1) (f a) := begin rw [deterministic], change measure.dirac (f a) s = s.indicator 1 (f a), @@ -340,7 +340,7 @@ begin end instance is_markov_kernel_deterministic {f : α → β} (hf : measurable f) : - is_markov_kernel (deterministic hf) := + is_markov_kernel (deterministic f hf) := ⟨λ a, by { rw deterministic_apply hf, apply_instance, }⟩ end deterministic diff --git a/src/probability/kernel/composition.lean b/src/probability/kernel/composition.lean index 569398e57c89b..011bc48df906c 100644 --- a/src/probability/kernel/composition.lean +++ b/src/probability/kernel/composition.lean @@ -662,7 +662,7 @@ begin end lemma deterministic_comp_eq_map (hf : measurable f) (κ : kernel α β) [is_s_finite_kernel κ] : - (deterministic hf ∘ₖ κ) = map κ f hf := + (deterministic f hf ∘ₖ κ) = map κ f hf := begin ext a s hs : 2, simp_rw [map_apply' _ _ _ hs, comp_apply _ _ _ hs, deterministic_apply' hf _ hs, @@ -670,7 +670,7 @@ begin end lemma comp_deterministic_eq_comap (κ : kernel α β) [is_s_finite_kernel κ] (hg : measurable g) : - (κ ∘ₖ deterministic hg) = comap κ g hg := + (κ ∘ₖ deterministic g hg) = comap κ g hg := begin ext a s hs : 2, simp_rw [comap_apply' _ _ _ s, comp_apply _ _ _ hs, deterministic_apply hg a, From 338fe44f54751b9f7deaca47ffca3509f53140ae Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Thu, 4 May 2023 11:37:58 +0000 Subject: [PATCH 0951/1291] feat(data/is_R_or_C/basic): drop 2 fields, golf (#18931) * Drop fields `inv_def_ax` and `div_I_ax` in `is_R_or_C`, deduce them instead. * Use lemmas about `ring_hom`s to prove properties of coercion, `is_R_or_C.re` etc. * Drop `is_R_or_C.of_real_hom` and `is_R_or_C.coe_hom`. * Drop `is_R_or_C.inv_zero` and `is_R_or_C.mul_inv_cancel`. * Move some lemmas to more appropriate sections. --- src/analysis/complex/basic.lean | 5 +- src/data/is_R_or_C/basic.lean | 193 ++++++++++++-------------------- 2 files changed, 73 insertions(+), 125 deletions(-) diff --git a/src/analysis/complex/basic.lean b/src/analysis/complex/basic.lean index 455d27e524793..4a046d4ec7e65 100644 --- a/src/analysis/complex/basic.lean +++ b/src/analysis/complex/basic.lean @@ -312,10 +312,7 @@ noncomputable instance : is_R_or_C ℂ := conj_I_ax := by simp only [complex.conj_I, ring_hom.coe_mk], norm_sq_eq_def_ax := λ z, by simp only [←complex.norm_sq_eq_abs, ←complex.norm_sq_apply, add_monoid_hom.coe_mk, complex.norm_eq_abs], - mul_im_I_ax := λ z, by simp only [mul_one, add_monoid_hom.coe_mk, complex.I_im], - inv_def_ax := λ z, by simp only [complex.inv_def, complex.norm_sq_eq_abs, complex.coe_algebra_map, - complex.of_real_eq_coe, complex.norm_eq_abs], - div_I_ax := complex.div_I } + mul_im_I_ax := λ z, by simp only [mul_one, add_monoid_hom.coe_mk, complex.I_im] } lemma _root_.is_R_or_C.re_eq_complex_re : ⇑(is_R_or_C.re : ℂ →+ ℝ) = complex.re := rfl lemma _root_.is_R_or_C.im_eq_complex_im : ⇑(is_R_or_C.im : ℂ →+ ℝ) = complex.im := rfl diff --git a/src/data/is_R_or_C/basic.lean b/src/data/is_R_or_C/basic.lean index dd036822025a4..c4c1b2924aa12 100644 --- a/src/data/is_R_or_C/basic.lean +++ b/src/data/is_R_or_C/basic.lean @@ -64,8 +64,6 @@ class is_R_or_C (K : Type*) (conj_I_ax : conj I = -I) (norm_sq_eq_def_ax : ∀ (z : K), ‖z‖^2 = (re z) * (re z) + (im z) * (im z)) (mul_im_I_ax : ∀ (z : K), (im z) * im I = im z) -(inv_def_ax : ∀ (z : K), z⁻¹ = conj z * 𝓚 ((‖z‖^2)⁻¹)) -(div_I_ax : ∀ (z : K), z / I = -(z * I)) end @@ -85,7 +83,7 @@ lemma of_real_alg (x : ℝ) : (x : K) = x • (1 : K) := algebra.algebra_map_eq_smul_one x lemma real_smul_eq_coe_mul (r : ℝ) (z : K) : r • z = (r : K) * z := -by rw [is_R_or_C.of_real_alg, ←smul_eq_mul, smul_assoc, smul_eq_mul, one_mul] +algebra.smul_def r z lemma real_smul_eq_coe_smul [add_comm_group E] [module K E] [module ℝ E] [is_scalar_tower ℝ K E] (r : ℝ) (x : E) : r • x = (r : K) • x := @@ -104,62 +102,77 @@ is_R_or_C.mul_re_ax @[simp, is_R_or_C_simps] lemma mul_im : ∀ z w : K, im (z * w) = re z * im w + im z * re w := is_R_or_C.mul_im_ax -theorem inv_def (z : K) : z⁻¹ = conj z * ((‖z‖^2)⁻¹:ℝ) := -is_R_or_C.inv_def_ax z - -theorem ext_iff : ∀ {z w : K}, z = w ↔ re z = re w ∧ im z = im w := -λ z w, { mp := by { rintro rfl, cc }, - mpr := by { rintro ⟨h₁,h₂⟩, rw [←re_add_im z, ←re_add_im w, h₁, h₂] } } - -theorem ext : ∀ {z w : K}, re z = re w → im z = im w → z = w := -by { simp_rw ext_iff, cc } +theorem ext_iff {z w : K} : z = w ↔ re z = re w ∧ im z = im w := +⟨λ h, h ▸ ⟨rfl, rfl⟩, λ ⟨h₁, h₂⟩, re_add_im z ▸ re_add_im w ▸ h₁ ▸ h₂ ▸ rfl⟩ +theorem ext {z w : K} (hre : re z = re w) (him : im z = im w) : z = w := +ext_iff.2 ⟨hre, him⟩ @[norm_cast] lemma of_real_zero : ((0 : ℝ) : K) = 0 := by rw [of_real_alg, zero_smul] @[simp, is_R_or_C_simps] lemma zero_re' : re (0 : K) = (0 : ℝ) := re.map_zero -@[norm_cast] lemma of_real_one : ((1 : ℝ) : K) = 1 := -by rw [of_real_alg, one_smul] -@[simp, is_R_or_C_simps] lemma one_re : re (1 : K) = 1 := by rw [←of_real_one, of_real_re] -@[simp, is_R_or_C_simps] lemma one_im : im (1 : K) = 0 := by rw [←of_real_one, of_real_im] +@[norm_cast] lemma of_real_one : ((1 : ℝ) : K) = 1 := map_one (algebra_map ℝ K) +@[simp, is_R_or_C_simps] lemma one_re : re (1 : K) = 1 := by rw [← of_real_one, of_real_re] +@[simp, is_R_or_C_simps] lemma one_im : im (1 : K) = 0 := by rw [← of_real_one, of_real_im] + +theorem of_real_injective : function.injective (coe : ℝ → K) := (algebra_map ℝ K).injective @[norm_cast] theorem of_real_inj {z w : ℝ} : (z : K) = (w : K) ↔ z = w := -{ mp := λ h, by { convert congr_arg re h; simp only [of_real_re] }, - mpr := λ h, by rw h } +algebra_map.coe_inj -@[simp, is_R_or_C_simps] lemma bit0_re (z : K) : re (bit0 z) = bit0 (re z) := -by simp only [bit0, map_add] +@[simp, is_R_or_C_simps] lemma bit0_re (z : K) : re (bit0 z) = bit0 (re z) := map_bit0 _ _ @[simp, is_R_or_C_simps] lemma bit1_re (z : K) : re (bit1 z) = bit1 (re z) := -by simp only [bit1, add_monoid_hom.map_add, bit0_re, add_right_inj, one_re] -@[simp, is_R_or_C_simps] lemma bit0_im (z : K) : im (bit0 z) = bit0 (im z) := -by simp only [bit0, map_add] +by simp only [bit1, map_add, bit0_re, one_re] +@[simp, is_R_or_C_simps] lemma bit0_im (z : K) : im (bit0 z) = bit0 (im z) := map_bit0 _ _ @[simp, is_R_or_C_simps] lemma bit1_im (z : K) : im (bit1 z) = bit0 (im z) := -by simp only [bit1, add_right_eq_self, add_monoid_hom.map_add, bit0_im, one_im] +by simp only [bit1, map_add, bit0_im, one_im, add_zero] -theorem of_real_eq_zero {z : ℝ} : (z : K) = 0 ↔ z = 0 := -by rw [←of_real_zero]; exact of_real_inj - -theorem of_real_ne_zero {z : ℝ} : (z : K) ≠ 0 ↔ z ≠ 0 := of_real_eq_zero.not +theorem of_real_eq_zero {x : ℝ} : (x : K) = 0 ↔ x = 0 := algebra_map.lift_map_eq_zero_iff x +theorem of_real_ne_zero {x : ℝ} : (x : K) ≠ 0 ↔ x ≠ 0 := of_real_eq_zero.not @[simp, is_R_or_C_simps, norm_cast, priority 900] -lemma of_real_add ⦃r s : ℝ⦄ : ((r + s : ℝ) : K) = r + s := -by { apply (@is_R_or_C.ext_iff K _ ((r + s : ℝ) : K) (r + s)).mpr, simp } +lemma of_real_add (r s : ℝ) : ((r + s : ℝ) : K) = r + s := algebra_map.coe_add _ _ @[simp, is_R_or_C_simps, norm_cast, priority 900] -lemma of_real_bit0 (r : ℝ) : ((bit0 r : ℝ) : K) = bit0 (r : K) := -ext_iff.2 $ by simp [bit0] +lemma of_real_bit0 (r : ℝ) : ((bit0 r : ℝ) : K) = bit0 (r : K) := of_real_add _ _ @[simp, is_R_or_C_simps, norm_cast, priority 900] lemma of_real_bit1 (r : ℝ) : ((bit1 r : ℝ) : K) = bit1 (r : K) := -ext_iff.2 $ by simp [bit1] +map_bit1 (algebra_map ℝ K) r + +@[simp, norm_cast, is_R_or_C_simps, priority 900] +lemma of_real_neg (r : ℝ) : ((-r : ℝ) : K) = -r := algebra_map.coe_neg r + +@[simp, norm_cast, is_R_or_C_simps, priority 900] +lemma of_real_sub (r s : ℝ) : ((r - s : ℝ) : K) = r - s := map_sub (algebra_map ℝ K) r s + +@[simp, is_R_or_C_simps, norm_cast, priority 900] +lemma of_real_sum {α : Type*} (s : finset α) (f : α → ℝ) : + ((∑ i in s, f i : ℝ) : K) = ∑ i in s, (f i : K) := +map_sum (algebra_map ℝ K) _ _ + +@[simp, is_R_or_C_simps, norm_cast] lemma of_real_finsupp_sum + {α M : Type*} [has_zero M] (f : α →₀ M) (g : α → M → ℝ) : + ((f.sum (λ a b, g a b) : ℝ) : K) = f.sum (λ a b, ((g a b) : K)) := +map_finsupp_sum (algebra_map ℝ K) f g @[simp, norm_cast, is_R_or_C_simps, priority 900] -lemma of_real_neg (r : ℝ) : ((-r : ℝ) : K) = -r := ext_iff.2 $ by simp +lemma of_real_mul (r s : ℝ) : ((r * s : ℝ) : K) = r * s := algebra_map.coe_mul _ _ @[simp, norm_cast, is_R_or_C_simps, priority 900] -lemma of_real_mul (r s : ℝ) : ((r * s : ℝ) : K) = r * s := ext_iff.2 $ by simp with is_R_or_C_simps +lemma of_real_pow (r : ℝ) (n : ℕ) : ((r ^ n : ℝ) : K) = r ^ n := map_pow (algebra_map ℝ K) r n + +@[simp, is_R_or_C_simps, norm_cast, priority 900] +lemma of_real_prod {α : Type*} (s : finset α) (f : α → ℝ) : + ((∏ i in s, f i : ℝ) : K) = ∏ i in s, (f i : K) := +ring_hom.map_prod _ _ _ + +@[simp, is_R_or_C_simps, norm_cast] lemma of_real_finsupp_prod + {α M : Type*} [has_zero M] (f : α →₀ M) (g : α → M → ℝ) : + ((f.prod (λ a b, g a b) : ℝ) : K) = f.prod (λ a b, ((g a b) : K)) := +ring_hom.map_finsupp_prod _ f g @[simp, norm_cast, is_R_or_C_simps] lemma of_real_smul (r x : ℝ) : r • (x : K) = (r : K) * (x : K) := @@ -201,14 +214,11 @@ lemma I_mul_I : (I : K) = 0 ∨ (I : K) * I = -1 := I_mul_I_ax by { rw ext_iff, simp only [of_real_im, conj_im, eq_self_iff_true, conj_re, and_self, neg_zero] } -@[simp, is_R_or_C_simps] lemma conj_bit0 (z : K) : conj (bit0 z) = bit0 (conj z) := -by simp only [bit0, ring_hom.map_add, eq_self_iff_true] -@[simp, is_R_or_C_simps] lemma conj_bit1 (z : K) : conj (bit1 z) = bit1 (conj z) := -by simp only [bit0, ext_iff, bit1_re, conj_im, eq_self_iff_true, conj_re, neg_add_rev, - and_self, bit1_im] +@[simp, is_R_or_C_simps] lemma conj_bit0 (z : K) : conj (bit0 z) = bit0 (conj z) := map_bit0 _ _ +@[simp, is_R_or_C_simps] lemma conj_bit1 (z : K) : conj (bit1 z) = bit1 (conj z) := map_bit1 _ _ @[simp, is_R_or_C_simps] lemma conj_neg_I : conj (-I) = (I : K) := -by simp only [conj_I, ring_hom.map_neg, eq_self_iff_true, neg_neg] +by rw [map_neg, conj_I, neg_neg] lemma conj_eq_re_sub_im (z : K) : conj z = re z - (im z) * I := by { rw ext_iff, simp only [add_zero, I_re, of_real_im, I_im, zero_sub, zero_mul, conj_im, @@ -305,24 +315,6 @@ theorem add_conj (z : K) : z + conj z = 2 * (re z) := by simp only [ext_iff, two_mul, map_add, add_zero, of_real_im, conj_im, of_real_re, eq_self_iff_true, add_right_neg, conj_re, and_self] -/-- The pseudo-coercion `of_real` as a `ring_hom`. -/ -noncomputable def of_real_hom : ℝ →+* K := algebra_map ℝ K - -/-- The coercion from reals as a `ring_hom`. -/ -noncomputable def coe_hom : ℝ →+* K := ⟨coe, of_real_one, of_real_mul, of_real_zero, of_real_add⟩ - -@[simp, norm_cast, is_R_or_C_simps, priority 900] lemma of_real_sub (r s : ℝ) : - ((r - s : ℝ) : K) = r - s := -ext_iff.2 $ by simp only [of_real_im, of_real_re, eq_self_iff_true, sub_zero, and_self, map_sub] - -@[simp, norm_cast, is_R_or_C_simps, priority 900] lemma of_real_pow (r : ℝ) (n : ℕ) : - ((r ^ n : ℝ) : K) = r ^ n := -begin - induction n, - { simp only [of_real_one, pow_zero]}, - { simp only [*, of_real_mul, pow_succ]} -end - theorem sub_conj (z : K) : z - conj z = (2 * im z) * I := by simp only [ext_iff, two_mul, sub_eq_add_neg, add_mul, map_add, add_zero, add_left_inj, zero_mul, map_add_neg, eq_self_iff_true, add_right_neg, and_self, neg_neg, mul_zero, neg_zero] @@ -333,42 +325,32 @@ by simp only [norm_sq_add, sub_eq_add_neg, ring_equiv.map_neg, mul_neg, norm_sq_neg, map_neg] lemma sqrt_norm_sq_eq_norm {z : K} : real.sqrt (norm_sq z) = ‖z‖ := -begin - have h₂ : ‖z‖ = real.sqrt (‖z‖^2) := (real.sqrt_sq (norm_nonneg z)).symm, - rw [h₂], - exact congr_arg real.sqrt (norm_sq_eq_def' z) -end +by rw [norm_sq_eq_def', real.sqrt_sq (norm_nonneg _)] /-! ### Inversion -/ -@[simp, is_R_or_C_simps] lemma inv_re (z : K) : re (z⁻¹) = re z / norm_sq z := -by simp only [inv_def, norm_sq_eq_def, norm_sq, division_def, - monoid_with_zero_hom.coe_mk, sub_zero, mul_zero] with is_R_or_C_simps -@[simp, is_R_or_C_simps] lemma inv_im (z : K) : im (z⁻¹) = im (-z) / norm_sq z := -by simp only [inv_def, norm_sq_eq_def, norm_sq, division_def, of_real_im, - monoid_with_zero_hom.coe_mk, of_real_re, zero_add, map_neg, mul_zero] - with is_R_or_C_simps - @[simp, norm_cast, is_R_or_C_simps, priority 900] -lemma of_real_inv (r : ℝ) : ((r⁻¹ : ℝ) : K) = r⁻¹ := +lemma of_real_inv (r : ℝ) : ((r⁻¹ : ℝ) : K) = r⁻¹ := map_inv₀ (algebra_map ℝ K) r + +theorem inv_def (z : K) : z⁻¹ = conj z * ((‖z‖^2)⁻¹:ℝ) := begin - rw ext_iff, - by_cases r = 0, - { simp only [h, of_real_zero, inv_zero, and_self, map_zero]}, - { simp only with is_R_or_C_simps, - field_simp [h, norm_sq] } + rcases eq_or_ne z 0 with (rfl | h₀), + { simp }, + { apply inv_eq_of_mul_eq_one_right, + rw [← mul_assoc, mul_conj, of_real_inv, ← norm_sq_eq_def', mul_inv_cancel], + rwa [of_real_ne_zero, ne.def, norm_sq_eq_zero] } end -protected lemma inv_zero : (0⁻¹ : K) = 0 := -by rw [← of_real_zero, ← of_real_inv, inv_zero] +@[simp, is_R_or_C_simps] lemma inv_re (z : K) : re (z⁻¹) = re z / norm_sq z := +by rw [inv_def, norm_sq_eq_def', mul_comm, of_real_mul_re, conj_re, div_eq_inv_mul] -protected theorem mul_inv_cancel {z : K} (h : z ≠ 0) : z * z⁻¹ = 1 := -by rw [inv_def, ←mul_assoc, mul_conj, ←of_real_mul, ←norm_sq_eq_def', - mul_inv_cancel (mt norm_sq_eq_zero.1 h), of_real_one] +@[simp, is_R_or_C_simps] lemma inv_im (z : K) : im (z⁻¹) = -im z / norm_sq z := +by rw [inv_def, norm_sq_eq_def', mul_comm, of_real_mul_im, conj_im, div_eq_inv_mul] lemma div_re (z w : K) : re (z / w) = re z * re w / norm_sq w + im z * im w / norm_sq w := by simp only [div_eq_mul_inv, mul_assoc, sub_eq_add_neg, neg_mul, mul_neg, neg_neg, map_neg] with is_R_or_C_simps + lemma div_im (z w : K) : im (z / w) = im z * re w / norm_sq w - re z * im w / norm_sq w := by simp only [div_eq_mul_inv, mul_assoc, sub_eq_add_neg, add_comm, neg_mul, mul_neg, map_neg] with is_R_or_C_simps @@ -378,23 +360,16 @@ lemma conj_inv (x : K) : conj (x⁻¹) = (conj x)⁻¹ := star_inv' _ @[simp, norm_cast, is_R_or_C_simps, priority 900] lemma of_real_div (r s : ℝ) : ((r / s : ℝ) : K) = r / s := -map_div₀ (@is_R_or_C.coe_hom K _) r s +map_div₀ (algebra_map ℝ K) r s lemma div_re_of_real {z : K} {r : ℝ} : re (z / r) = re z / r := -begin - by_cases h : r = 0, - { simp only [h, of_real_zero, div_zero, zero_re']}, - { change r ≠ 0 at h, - rw [div_eq_mul_inv, ←of_real_inv, div_eq_mul_inv], - simp only [one_div, of_real_im, of_real_re, sub_zero, mul_re, mul_zero]} -end +by rw [div_eq_inv_mul, div_eq_inv_mul, ← of_real_inv, of_real_mul_re] @[simp, norm_cast, is_R_or_C_simps, priority 900] lemma of_real_zpow (r : ℝ) (n : ℤ) : ((r ^ n : ℝ) : K) = r ^ n := -map_zpow₀ (@is_R_or_C.coe_hom K _) r n +map_zpow₀ (algebra_map ℝ K) r n -lemma I_mul_I_of_nonzero : (I : K) ≠ 0 → (I : K) * I = -1 := -by { have := I_mul_I_ax, tauto } +lemma I_mul_I_of_nonzero : (I : K) ≠ 0 → (I : K) * I = -1 := I_mul_I_ax.resolve_left @[simp, is_R_or_C_simps] lemma div_I (z : K) : z / I = -(z * I) := begin @@ -422,7 +397,7 @@ by simp only [←sqrt_norm_sq_eq_norm, norm_sq_conj] @[simp, is_R_or_C_simps, norm_cast, priority 900] theorem of_real_nat_cast (n : ℕ) : ((n : ℝ) : K) = n := -map_nat_cast (@of_real_hom K _) n +map_nat_cast (algebra_map ℝ K) n @[simp, is_R_or_C_simps, norm_cast] lemma nat_cast_re (n : ℕ) : re (n : K) = n := by rw [← of_real_nat_cast, of_real_re] @@ -431,7 +406,7 @@ by rw [← of_real_nat_cast, of_real_re] by rw [← of_real_nat_cast, of_real_im] @[simp, is_R_or_C_simps, norm_cast, priority 900] -lemma of_real_int_cast (n : ℤ) : ((n : ℝ) : K) = n := map_int_cast (@of_real_hom K _) n +lemma of_real_int_cast (n : ℤ) : ((n : ℝ) : K) = n := map_int_cast (algebra_map ℝ K) n @[simp, is_R_or_C_simps, norm_cast] lemma int_cast_re (n : ℤ) : re (n : K) = n := by rw [← of_real_int_cast, of_real_re] @@ -441,7 +416,7 @@ by rw [← of_real_int_cast, of_real_im] @[simp, is_R_or_C_simps, norm_cast, priority 900] theorem of_real_rat_cast (n : ℚ) : ((n : ℝ) : K) = n := -map_rat_cast (@is_R_or_C.of_real_hom K _) n +map_rat_cast (algebra_map ℝ K) n @[simp, is_R_or_C_simps, norm_cast] lemma rat_cast_re (q : ℚ) : re (q : K) = q := by rw [← of_real_rat_cast, of_real_re] @@ -646,26 +621,6 @@ lemma is_cau_seq_abs {f : ℕ → K} (hf : is_cau_seq abs f) : λ ε ε0, let ⟨i, hi⟩ := hf ε ε0 in ⟨i, λ j hj, lt_of_le_of_lt (abs_abs_sub_le_abs_sub _ _) (hi j hj)⟩ -@[simp, is_R_or_C_simps, norm_cast, priority 900] -lemma of_real_prod {α : Type*} (s : finset α) (f : α → ℝ) : - ((∏ i in s, f i : ℝ) : K) = ∏ i in s, (f i : K) := -ring_hom.map_prod _ _ _ - -@[simp, is_R_or_C_simps, norm_cast, priority 900] -lemma of_real_sum {α : Type*} (s : finset α) (f : α → ℝ) : - ((∑ i in s, f i : ℝ) : K) = ∑ i in s, (f i : K) := -ring_hom.map_sum _ _ _ - -@[simp, is_R_or_C_simps, norm_cast] lemma of_real_finsupp_sum - {α M : Type*} [has_zero M] (f : α →₀ M) (g : α → M → ℝ) : - ((f.sum (λ a b, g a b) : ℝ) : K) = f.sum (λ a b, ((g a b) : K)) := -ring_hom.map_finsupp_sum _ f g - -@[simp, is_R_or_C_simps, norm_cast] lemma of_real_finsupp_prod - {α M : Type*} [has_zero M] (f : α →₀ M) (g : α → M → ℝ) : - ((f.prod (λ a b, g a b) : ℝ) : K) = f.prod (λ a b, ((g a b) : K)) := -ring_hom.map_finsupp_prod _ f g - end is_R_or_C @@ -690,10 +645,6 @@ noncomputable instance real.is_R_or_C : is_R_or_C ℝ := norm_sq_eq_def_ax := λ z, by simp only [sq, real.norm_eq_abs, ←abs_mul, abs_mul_self z, add_zero, mul_zero, add_monoid_hom.zero_apply, add_monoid_hom.id_apply], mul_im_I_ax := λ z, by simp only [mul_zero, add_monoid_hom.zero_apply], - inv_def_ax := λ z, by simp only [star_ring_end_apply, star, sq, real.norm_eq_abs, - abs_mul_abs_self, ←div_eq_mul_inv, algebra.id.map_eq_id, id.def, ring_hom.id_apply, - div_self_mul_self'], - div_I_ax := λ z, by simp only [div_zero, mul_zero, neg_zero], .. real.densely_normed_field, .. real.metric_space } end instances From 2f39bcbc98f8255490f8d4562762c9467694c809 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 4 May 2023 11:38:00 +0000 Subject: [PATCH 0952/1291] chore(ring_theory/ideal/quotient): add missing smul compatibility instances (#18934) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The low priority is needed to avoid a timeout in what will be `ring_theory/kaehler.lean`. These instances are more general cases of the instances implied by `algebra α c.quotient` and `algebra α (R ⧸ I)`, which work for cases like `α = units R`. --- src/ring_theory/congruence.lean | 17 +++++++++++++++++ src/ring_theory/ideal/quotient.lean | 14 ++++++++++++++ 2 files changed, 31 insertions(+) diff --git a/src/ring_theory/congruence.lean b/src/ring_theory/congruence.lean index f0c339ea28545..865204133e71c 100644 --- a/src/ring_theory/congruence.lean +++ b/src/ring_theory/congruence.lean @@ -247,6 +247,23 @@ function.surjective.comm_ring _ quotient.surjective_quotient_mk' rfl rfl (λ _ _, rfl) (λ _ _, rfl) (λ _, rfl) (λ _ _, rfl) (λ _ _, rfl) (λ _ _, rfl) (λ _ _, rfl) (λ _, rfl) (λ _, rfl) +instance is_scalar_tower_right [has_add R] [mul_one_class R] [has_smul α R] [is_scalar_tower α R R] + (c : ring_con R) : + is_scalar_tower α c.quotient c.quotient := +{ smul_assoc := λ a, quotient.ind₂' $ by exact λ m₁ m₂, + congr_arg quotient.mk' $ smul_mul_assoc _ _ _ } + +instance smul_comm_class [has_add R] [mul_one_class R] [has_smul α R] + [is_scalar_tower α R R] [smul_comm_class α R R] (c : ring_con R) : + smul_comm_class α c.quotient c.quotient := +{ smul_comm := λ a, quotient.ind₂' $ by exact λ m₁ m₂, + congr_arg quotient.mk' $ (mul_smul_comm _ _ _).symm } + +instance smul_comm_class' [has_add R] [mul_one_class R] [has_smul α R] + [is_scalar_tower α R R] [smul_comm_class R α R] (c : ring_con R) : + smul_comm_class c.quotient α c.quotient := +by haveI := smul_comm_class.symm R α R; exact smul_comm_class.symm _ _ _ + instance [monoid α] [non_assoc_semiring R] [distrib_mul_action α R] [is_scalar_tower α R R] (c : ring_con R) : distrib_mul_action α c.quotient := diff --git a/src/ring_theory/ideal/quotient.lean b/src/ring_theory/ideal/quotient.lean index cda8ece37d175..4cefd306702b4 100644 --- a/src/ring_theory/ideal/quotient.lean +++ b/src/ring_theory/ideal/quotient.lean @@ -72,6 +72,20 @@ instance comm_ring (I : ideal R) : comm_ring (R ⧸ I) := { ..submodule.quotient.add_comm_group I, -- to help with unification ..(quotient.ring_con I)^.quotient.comm_ring } +-- this instance is harder to find than the one via `algebra α (R ⧸ I)`, so use a lower priority +@[priority 100] +instance is_scalar_tower_right {α} [has_smul α R] [is_scalar_tower α R R] : + is_scalar_tower α (R ⧸ I) (R ⧸ I) := +(quotient.ring_con I)^.is_scalar_tower_right + +instance smul_comm_class {α} [has_smul α R] [is_scalar_tower α R R] [smul_comm_class α R R] : + smul_comm_class α (R ⧸ I) (R ⧸ I) := +(quotient.ring_con I)^.smul_comm_class + +instance smul_comm_class' {α} [has_smul α R] [is_scalar_tower α R R] [smul_comm_class R α R] : + smul_comm_class (R ⧸ I) α (R ⧸ I) := +(quotient.ring_con I)^.smul_comm_class' + /-- The ring homomorphism from a ring `R` to a quotient ring `R/I`. -/ def mk (I : ideal R) : R →+* (R ⧸ I) := ⟨λ a, submodule.quotient.mk a, rfl, λ _ _, rfl, rfl, λ _ _, rfl⟩ From 26ae6f6771ba2d8c76ac47efb84ef2908f178630 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 4 May 2023 11:38:01 +0000 Subject: [PATCH 0953/1291] chore(ring_theory/derivation): generalize tower instance (#18937) --- src/ring_theory/derivation.lean | 14 +++++++++----- 1 file changed, 9 insertions(+), 5 deletions(-) diff --git a/src/ring_theory/derivation.lean b/src/ring_theory/derivation.lean index 5d70b42f02def..8816fcedc40de 100644 --- a/src/ring_theory/derivation.lean +++ b/src/ring_theory/derivation.lean @@ -159,8 +159,9 @@ instance : inhabited (derivation R A M) := ⟨0⟩ section scalar -variables {S : Type*} [monoid S] [distrib_mul_action S M] [smul_comm_class R S M] - [smul_comm_class S A M] +variables {S T : Type*} +variables [monoid S] [distrib_mul_action S M] [smul_comm_class R S M] [smul_comm_class S A M] +variables [monoid T] [distrib_mul_action T M] [smul_comm_class R T M] [smul_comm_class T A M] @[priority 100] instance : has_smul S (derivation R A M) := @@ -190,6 +191,12 @@ instance [distrib_mul_action Sᵐᵒᵖ M] [is_central_scalar S M] : is_central_scalar S (derivation R A M) := { op_smul_eq_smul := λ _ _, ext $ λ _, op_smul_eq_smul _ _} +instance [has_smul S T] [is_scalar_tower S T M] : is_scalar_tower S T (derivation R A M) := +⟨λ x y z, ext $ λ a, smul_assoc _ _ _⟩ + +instance [smul_comm_class S T M] : smul_comm_class S T (derivation R A M) := +⟨λ x y z, ext $ λ a, smul_comm _ _ _⟩ + end scalar @[priority 100] @@ -197,9 +204,6 @@ instance {S : Type*} [semiring S] [module S M] [smul_comm_class R S M] [smul_com module S (derivation R A M) := function.injective.module S coe_fn_add_monoid_hom coe_injective coe_smul -instance [is_scalar_tower R A M] : is_scalar_tower R A (derivation R A M) := -⟨λ x y z, ext (λ a, smul_assoc _ _ _)⟩ - section push_forward variables {N : Type*} [add_comm_monoid N] [module A N] [module R N] [is_scalar_tower R A M] From 570e9f4877079b3a923135b3027ac3be8695ab8c Mon Sep 17 00:00:00 2001 From: "Filippo A. E. Nuccio" Date: Thu, 4 May 2023 13:44:28 +0000 Subject: [PATCH 0954/1291] feat(ring_theory/localization/away) : Add `num_denom` section (#18830) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Added a section `num_denom`: the main result is the lemma `exists_reduced_fraction` that shows that every non-zero element `b` in a `localization.away x` of a UFM can be written in a unique way as `b=x^n * a` with `n : ℤ` and `a` not divisible by `x`. Co-authored-by: Vierkantor --- src/ring_theory/localization/away.lean | 142 ++++++++++++++++++ .../unique_factorization_domain.lean | 28 +++- 2 files changed, 162 insertions(+), 8 deletions(-) diff --git a/src/ring_theory/localization/away.lean b/src/ring_theory/localization/away.lean index 39d04cb19893d..0ea25713791eb 100644 --- a/src/ring_theory/localization/away.lean +++ b/src/ring_theory/localization/away.lean @@ -13,6 +13,8 @@ import ring_theory.localization.basic * `is_localization.away (x : R) S` expresses that `S` is a localization away from `x`, as an abbreviation of `is_localization (submonoid.powers x) S` + * `exists_reduced_fraction (hb : b ≠ 0)` produces a reduced fraction of the form `b = a * x^n` for + some `n : ℤ` and some `a : R` that is not divisible by `x`. ## Implementation notes @@ -194,3 +196,143 @@ lemma is_localization.away.finite_presentation (r : R) {S} [comm_ring S] [algebr [is_localization.away r S] : algebra.finite_presentation R S := (adjoin_root.finite_presentation _).equiv $ (localization.away_equiv_adjoin r).symm.trans $ is_localization.alg_equiv (submonoid.powers r) _ _ + +section num_denom + +open multiplicity unique_factorization_monoid is_localization + +variable (x : R) + +variables (B : Type*) [comm_ring B] [algebra R B] [is_localization.away x B] + +/-- `self_zpow x (m : ℤ)` is `x ^ m` as an element of the localization away from `x`. -/ +noncomputable def self_zpow (m : ℤ) : B := +if hm : 0 ≤ m +then algebra_map _ _ x ^ m.nat_abs +else mk' _ (1 : R) (submonoid.pow x m.nat_abs) + +lemma self_zpow_of_nonneg {n : ℤ} (hn : 0 ≤ n) : self_zpow x B n = + algebra_map R B x ^ n.nat_abs := +dif_pos hn + +@[simp] lemma self_zpow_coe_nat (d : ℕ) : self_zpow x B d = (algebra_map R B x)^d := +self_zpow_of_nonneg _ _ (int.coe_nat_nonneg d) + +@[simp] lemma self_zpow_zero : self_zpow x B 0 = 1 := +by simp [self_zpow_of_nonneg _ _ le_rfl] + +lemma self_zpow_of_neg {n : ℤ} (hn : n < 0) : + self_zpow x B n = mk' _ (1 : R) (submonoid.pow x n.nat_abs) := +dif_neg hn.not_le + +lemma self_zpow_of_nonpos {n : ℤ} (hn : n ≤ 0) : + self_zpow x B n = mk' _ (1 : R) (submonoid.pow x n.nat_abs) := +begin + by_cases hn0 : n = 0, + { simp [hn0, self_zpow_zero, submonoid.pow_apply] }, + { simp [self_zpow_of_neg _ _ (lt_of_le_of_ne hn hn0)] } +end + +@[simp] lemma self_zpow_neg_coe_nat (d : ℕ) : + self_zpow x B (-d) = mk' _ (1 : R) (submonoid.pow x d) := +by simp [self_zpow_of_nonpos _ _ (neg_nonpos.mpr (int.coe_nat_nonneg d))] + +@[simp] lemma self_zpow_sub_cast_nat {n m : ℕ} : + self_zpow x B (n - m) = mk' _ (x ^ n) (submonoid.pow x m) := +begin + by_cases h : m ≤ n, + { rw [is_localization.eq_mk'_iff_mul_eq, submonoid.pow_apply, subtype.coe_mk, + ← int.coe_nat_sub h, self_zpow_coe_nat, ← map_pow, ← map_mul, ← pow_add, + nat.sub_add_cancel h] }, + { rw [← neg_sub, ← int.coe_nat_sub (le_of_not_le h), self_zpow_neg_coe_nat, + is_localization.mk'_eq_iff_eq], + simp [submonoid.pow_apply, ← pow_add, nat.sub_add_cancel (le_of_not_le h)] } +end + +@[simp] lemma self_zpow_add {n m : ℤ} : + self_zpow x B (n + m) = self_zpow x B n * self_zpow x B m := +begin + cases le_or_lt 0 n with hn hn; cases le_or_lt 0 m with hm hm, + { rw [self_zpow_of_nonneg _ _ hn, self_zpow_of_nonneg _ _ hm, + self_zpow_of_nonneg _ _ (add_nonneg hn hm), int.nat_abs_add_nonneg hn hm, pow_add] }, + { have : n + m = n.nat_abs - m.nat_abs, + { rw [int.nat_abs_of_nonneg hn, int.of_nat_nat_abs_of_nonpos hm.le, sub_neg_eq_add] }, + rw [self_zpow_of_nonneg _ _ hn, self_zpow_of_neg _ _ hm, + this, self_zpow_sub_cast_nat, is_localization.mk'_eq_mul_mk'_one, map_pow] }, + { have : n + m = m.nat_abs - n.nat_abs, + { rw [int.nat_abs_of_nonneg hm, int.of_nat_nat_abs_of_nonpos hn.le, sub_neg_eq_add, add_comm] }, + rw [self_zpow_of_nonneg _ _ hm, self_zpow_of_neg _ _ hn, + this, self_zpow_sub_cast_nat, is_localization.mk'_eq_mul_mk'_one, map_pow, mul_comm] }, + { rw [self_zpow_of_neg _ _ hn, self_zpow_of_neg _ _ hm, self_zpow_of_neg _ _ (add_neg hn hm), + int.nat_abs_add_neg hn hm, ← mk'_mul, one_mul], + congr, + ext, + simp [pow_add] }, +end + +lemma self_zpow_mul_neg (d : ℤ) : self_zpow x B d * self_zpow x B (-d) = 1 := +begin + by_cases hd : d ≤ 0, + { erw [self_zpow_of_nonpos x B hd, self_zpow_of_nonneg, ← map_pow, int.nat_abs_neg, + is_localization.mk'_spec, map_one], + apply nonneg_of_neg_nonpos, + rwa [neg_neg]}, + { erw [self_zpow_of_nonneg x B (le_of_not_le hd), self_zpow_of_nonpos, ← map_pow, int.nat_abs_neg, + @is_localization.mk'_spec' R _ (submonoid.powers x) B _ _ _ 1 (submonoid.pow x d.nat_abs), + map_one], + refine nonpos_of_neg_nonneg (le_of_lt _), + rwa [neg_neg, ← not_le] }, +end + +lemma self_zpow_neg_mul (d : ℤ) : self_zpow x B (-d) * self_zpow x B d = 1 := +by rw [mul_comm, self_zpow_mul_neg x B d] + + +lemma self_zpow_pow_sub (a : R) (b : B) (m d : ℤ) : + (self_zpow x B (m - d)) * mk' B a (1 : submonoid.powers x) = b ↔ + (self_zpow x B m) * mk' B a (1 : submonoid.powers x) = (self_zpow x B d) * b := +begin + rw [sub_eq_add_neg, self_zpow_add, mul_assoc, mul_comm _ (mk' B a 1), ← mul_assoc], + split, + { intro h, + have := congr_arg (λ s : B, s * self_zpow x B d) h, + simp only at this, + rwa [mul_assoc, mul_assoc, self_zpow_neg_mul, mul_one, mul_comm b _] at this}, + { intro h, + have := congr_arg (λ s : B, s * self_zpow x B (-d)) h, + simp only at this, + rwa [mul_comm _ b, mul_assoc b _ _, self_zpow_mul_neg, mul_one] at this} +end + + +variables [is_domain R] [normalization_monoid R] [unique_factorization_monoid R] + + +theorem exists_reduced_fraction' {b : B} (hb : b ≠ 0) (hx : irreducible x) : + ∃ (a : R) (n : ℤ), ¬ x ∣ a ∧ + self_zpow x B n * algebra_map R B a = b := +begin + classical, + obtain ⟨⟨a₀, y⟩, H⟩ := surj (submonoid.powers x) b, + obtain ⟨d, hy⟩ := (submonoid.mem_powers_iff y.1 x).mp y.2, + have ha₀ : a₀ ≠ 0, + { haveI := @is_domain_of_le_non_zero_divisors B _ R _ _ _ (submonoid.powers x) _ + (powers_le_non_zero_divisors_of_no_zero_divisors hx.ne_zero), + simp only [map_zero, ← subtype.val_eq_coe, ← hy, map_pow] at H, + apply ((injective_iff_map_eq_zero' (algebra_map R B)).mp _ a₀).mpr.mt, + rw ← H, + apply mul_ne_zero hb (pow_ne_zero _ _), + exact is_localization.to_map_ne_zero_of_mem_non_zero_divisors B + (powers_le_non_zero_divisors_of_no_zero_divisors (hx.ne_zero)) + (mem_non_zero_divisors_iff_ne_zero.mpr hx.ne_zero), + exact is_localization.injective B (powers_le_non_zero_divisors_of_no_zero_divisors + (hx.ne_zero)) }, + simp only [← subtype.val_eq_coe, ← hy] at H, + obtain ⟨m, a, hyp1, hyp2⟩ := max_power_factor ha₀ hx, + refine ⟨a, m-d, _⟩, + rw [← mk'_one B, self_zpow_pow_sub, self_zpow_coe_nat, self_zpow_coe_nat, ← map_pow _ _ d, + mul_comm _ b, H, hyp2, map_mul, map_pow _ _ m], + exact ⟨hyp1, (congr_arg _ (is_localization.mk'_one _ _))⟩, +end + +end num_denom diff --git a/src/ring_theory/unique_factorization_domain.lean b/src/ring_theory/unique_factorization_domain.lean index 471af7ef520d2..d4c443fe63a91 100644 --- a/src/ring_theory/unique_factorization_domain.lean +++ b/src/ring_theory/unique_factorization_domain.lean @@ -863,12 +863,12 @@ lemma pow_eq_pow_iff {a : R} (ha0 : a ≠ 0) (ha1 : ¬ is_unit a) {i j : ℕ} : (pow_right_injective ha0 ha1).eq_iff section multiplicity -variables [nontrivial R] [normalization_monoid R] [decidable_eq R] +variables [nontrivial R] [normalization_monoid R] variables [dec_dvd : decidable_rel (has_dvd.dvd : R → R → Prop)] open multiplicity multiset include dec_dvd -lemma le_multiplicity_iff_replicate_le_normalized_factors {a b : R} {n : ℕ} +lemma le_multiplicity_iff_replicate_le_normalized_factors [decidable_eq R] {a b : R} {n : ℕ} (ha : irreducible a) (hb : b ≠ 0) : ↑n ≤ multiplicity a b ↔ replicate n (normalize a) ≤ normalized_factors b := begin @@ -894,8 +894,8 @@ the normalized factor occurs in the `normalized_factors`. See also `count_normalized_factors_eq` which expands the definition of `multiplicity` to produce a specification for `count (normalized_factors _) _`.. -/ -lemma multiplicity_eq_count_normalized_factors {a b : R} (ha : irreducible a) (hb : b ≠ 0) : - multiplicity a b = (normalized_factors b).count (normalize a) := +lemma multiplicity_eq_count_normalized_factors [decidable_eq R] {a b : R} (ha : irreducible a) + (hb : b ≠ 0) : multiplicity a b = (normalized_factors b).count (normalize a) := begin apply le_antisymm, { apply part_enat.le_of_lt_add_one, @@ -912,8 +912,8 @@ the number of times it divides `x`. See also `multiplicity_eq_count_normalized_factors` if `n` is given by `multiplicity p x`. -/ -lemma count_normalized_factors_eq {p x : R} (hp : irreducible p) (hnorm : normalize p = p) {n : ℕ} - (hle : p^n ∣ x) (hlt : ¬ (p^(n+1) ∣ x)) : +lemma count_normalized_factors_eq [decidable_eq R] {p x : R} (hp : irreducible p) + (hnorm : normalize p = p) {n : ℕ} (hle : p^n ∣ x) (hlt : ¬ (p^(n+1) ∣ x)) : (normalized_factors x).count p = n := begin letI : decidable_rel ((∣) : R → R → Prop) := λ _ _, classical.prop_decidable _, @@ -931,8 +931,8 @@ the number of times it divides `x`. This is a slightly more general version of See also `multiplicity_eq_count_normalized_factors` if `n` is given by `multiplicity p x`. -/ -lemma count_normalized_factors_eq' {p x : R} (hp : p = 0 ∨ irreducible p) (hnorm : normalize p = p) - {n : ℕ} (hle : p^n ∣ x) (hlt : ¬ (p^(n+1) ∣ x)) : +lemma count_normalized_factors_eq' [decidable_eq R] {p x : R} (hp : p = 0 ∨ irreducible p) + (hnorm : normalize p = p) {n : ℕ} (hle : p^n ∣ x) (hlt : ¬ (p^(n+1) ∣ x)) : (normalized_factors x).count p = n := begin rcases hp with rfl|hp, @@ -943,6 +943,18 @@ begin { exact count_normalized_factors_eq hp hnorm hle hlt } end +lemma max_power_factor {a₀ : R} {x : R} (h : a₀ ≠ 0) (hx : irreducible x) : + ∃ n : ℕ, ∃ a : R, ¬ x ∣ a ∧ a₀ = x ^ n * a := +begin + classical, + let n := (normalized_factors a₀).count (normalize x), + obtain ⟨a, ha1, ha2⟩ := (@exists_eq_pow_mul_and_not_dvd R _ _ x a₀ + (ne_top_iff_finite.mp (part_enat.ne_top_iff.mpr _))), + simp_rw [← (multiplicity_eq_count_normalized_factors hx h).symm] at ha1, + use [n, a, ha2, ha1], + use [n, (multiplicity_eq_count_normalized_factors hx h)], +end + end multiplicity section multiplicative From 738054fa93d43512da144ec45ce799d18fd44248 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 4 May 2023 20:00:45 +0000 Subject: [PATCH 0955/1291] chore(number_theory/modular_forms/slash_invariant_forms): golf and rename (#18933) In the `slash_action` namespace, this renames: * `right_action` to `slash_mul` (and reverses the direction), to match `slash_one` * `add_action` to `add_slash`, to match `zero_slash` * `smul_action` to `smul_slash`, to match `zero_slash` --- .../modular_forms/slash_actions.lean | 30 +++++++++---------- .../modular_forms/slash_invariant_forms.lean | 10 +++---- 2 files changed, 19 insertions(+), 21 deletions(-) diff --git a/src/number_theory/modular_forms/slash_actions.lean b/src/number_theory/modular_forms/slash_actions.lean index e03cf93f4925d..9c980afa75e78 100644 --- a/src/number_theory/modular_forms/slash_actions.lean +++ b/src/number_theory/modular_forms/slash_actions.lean @@ -39,9 +39,9 @@ class slash_action (β G α γ : Type*) [group G] [add_monoid α] [has_smul γ (map : β → G → α → α) (zero_slash : ∀ (k : β) (g : G), map k g 0 = 0) (slash_one : ∀ (k : β) (a : α) , map k 1 a = a) -(right_action : ∀ (k : β) (g h : G) (a : α), map k h (map k g a) = map k (g * h) a ) -(smul_action : ∀ (k : β) (g : G) (a : α) (z : γ), map k g (z • a) = z • (map k g a)) -(add_action : ∀ (k : β) (g : G) (a b : α), map k g (a + b) = map k g a + map k g b) +(slash_mul : ∀ (k : β) (g h : G) (a : α), map k (g * h) a =map k h (map k g a)) +(smul_slash : ∀ (k : β) (g : G) (a : α) (z : γ), map k g (z • a) = z • (map k g a)) +(add_slash : ∀ (k : β) (g : G) (a b : α), map k g (a + b) = map k g a + map k g b) localized "notation (name := modular_form.slash) f ` ∣[`:100 k `;` γ `] `:0 a :100 := slash_action.map γ k a f" in modular_form @@ -52,18 +52,18 @@ localized "notation (name := modular_form.slash_complex) f ` ∣[`:100 k `] `:0 @[simp] lemma slash_action.neg_slash {β G α γ : Type*} [group G] [add_group α] [has_smul γ α] [slash_action β G α γ] (k : β) (g : G) (a : α) : (-a) ∣[k;γ] g = - (a ∣[k;γ] g) := -eq_neg_of_add_eq_zero_left $ by rw [←slash_action.add_action, add_left_neg, slash_action.zero_slash] +eq_neg_of_add_eq_zero_left $ by rw [←slash_action.add_slash, add_left_neg, slash_action.zero_slash] @[simp] lemma slash_action.smul_slash_of_tower {R β G α : Type*} (γ : Type*) [group G] [add_group α] [monoid γ] [mul_action γ α] [has_smul R γ] [has_smul R α] [is_scalar_tower R γ α] [slash_action β G α γ] (k : β) (g : G) (a : α) (r : R) : (r • a) ∣[k;γ] g = r • (a ∣[k;γ] g) := -by rw [←smul_one_smul γ r a, slash_action.smul_action, smul_one_smul] +by rw [←smul_one_smul γ r a, slash_action.smul_slash, smul_one_smul] attribute [simp] slash_action.zero_slash slash_action.slash_one - slash_action.smul_action slash_action.add_action + slash_action.smul_slash slash_action.add_slash /--Slash_action induced by a monoid homomorphism.-/ def monoid_hom_slash_action {β G H α γ : Type*} [group G] [add_monoid α] [has_smul γ α] @@ -71,9 +71,9 @@ def monoid_hom_slash_action {β G H α γ : Type*} [group G] [add_monoid α] [ha { map := λ k g, slash_action.map γ k (h g), zero_slash := λ k g, slash_action.zero_slash k (h g), slash_one := λ k a, by simp only [map_one, slash_action.slash_one], - right_action := λ k g gg a, by simp only [map_mul, slash_action.right_action], - smul_action := λ _ _, slash_action.smul_action _ _, - add_action := λ _ g _ _, slash_action.add_action _ (h g) _ _,} + slash_mul := λ k g gg a, by simp only [map_mul, slash_action.slash_mul], + smul_slash := λ _ _, slash_action.smul_slash _ _, + add_slash := λ _ g _ _, slash_action.add_slash _ (h g) _ _,} namespace modular_form @@ -90,8 +90,8 @@ section -- temporary notation until the instance is built local notation f ` ∣[`:100 k `]`:0 γ :100 := modular_form.slash k γ f -private lemma slash_right_action (k : ℤ) (A B : GL(2, ℝ)⁺) (f : ℍ → ℂ) : - (f ∣[k] A) ∣[k] B = f ∣[k] (A * B) := +private lemma slash_mul (k : ℤ) (A B : GL(2, ℝ)⁺) (f : ℍ → ℂ) : + f ∣[k] (A * B) = (f ∣[k] A) ∣[k] B := begin ext1, simp_rw [slash,(upper_half_plane.denom_cocycle A B x)], @@ -109,7 +109,7 @@ begin simp_rw [this, ← mul_assoc, ←mul_zpow], end -private lemma slash_add (k : ℤ) (A : GL(2, ℝ)⁺) (f g : ℍ → ℂ) : +private lemma add_slash (k : ℤ) (A : GL(2, ℝ)⁺) (f g : ℍ → ℂ) : (f + g) ∣[k] A = (f ∣[k] A) + (g ∣[k] A) := begin ext1, @@ -140,9 +140,9 @@ instance : slash_action ℤ GL(2, ℝ)⁺ (ℍ → ℂ) ℂ := { map := slash, zero_slash := zero_slash, slash_one := slash_one, - right_action := slash_right_action, - smul_action := smul_slash, - add_action := slash_add } + slash_mul := slash_mul, + smul_slash := smul_slash, + add_slash := add_slash } end diff --git a/src/number_theory/modular_forms/slash_invariant_forms.lean b/src/number_theory/modular_forms/slash_invariant_forms.lean index 7ed80a3f81de8..98b6074ce71ee 100644 --- a/src/number_theory/modular_forms/slash_invariant_forms.lean +++ b/src/number_theory/modular_forms/slash_invariant_forms.lean @@ -85,7 +85,7 @@ instance slash_invariant_form_class.coe_to_fun [slash_invariant_form_class F Γ has_coe_to_fun F (λ _, ℍ → ℂ) := fun_like.has_coe_to_fun @[simp] lemma slash_action_eqn [slash_invariant_form_class F Γ k] (f : F) (γ : Γ) : - slash_action.map ℂ k γ ⇑f = ⇑f := slash_invariant_form_class.slash_action_eq f γ + ⇑f ∣[k] γ = ⇑f := slash_invariant_form_class.slash_action_eq f γ lemma slash_action_eqn' (k : ℤ) (Γ : subgroup SL(2, ℤ)) [slash_invariant_form_class F Γ k] (f : F) (γ : Γ) (z : ℍ) : f (γ • z) = ((↑ₘ[ℤ]γ 1 0 : ℂ) * z +(↑ₘ[ℤ]γ 1 1 : ℂ))^k * f z := @@ -103,7 +103,7 @@ instance [slash_invariant_form_class F Γ k] : has_coe_t F (slash_invariant_form instance has_add : has_add (slash_invariant_form Γ k) := ⟨ λ f g, { to_fun := f + g, - slash_action_eq' := λ γ, by convert slash_action.add_action k γ (f : ℍ → ℂ) g; simp } ⟩ + slash_action_eq' := λ γ, by rw [slash_action.add_slash, slash_action_eqn, slash_action_eqn] }⟩ @[simp] lemma coe_add (f g : slash_invariant_form Γ k) : ⇑(f + g) = f + g := rfl @[simp] lemma add_apply (f g : slash_invariant_form Γ k) (z : ℍ) : (f + g) z = f z + g z := rfl @@ -120,8 +120,7 @@ variables {α : Type*} [has_smul α ℂ] [is_scalar_tower α ℂ ℂ] instance has_smul : has_smul α (slash_invariant_form Γ k) := ⟨ λ c f, { to_fun := c • f, - slash_action_eq' := λ γ, by rw [←smul_one_smul ℂ c ⇑f, slash_action.smul_action k γ ⇑f, - slash_action_eqn] }⟩ + slash_action_eq' := λ γ, by rw [slash_action.smul_slash_of_tower, slash_action_eqn] }⟩ @[simp] lemma coe_smul (f : slash_invariant_form Γ k) (n : α) : ⇑(n • f) = n • f := rfl @[simp] lemma smul_apply (f : slash_invariant_form Γ k) (n : α) (z : ℍ) : @@ -132,8 +131,7 @@ end instance has_neg : has_neg (slash_invariant_form Γ k) := ⟨ λ f, { to_fun := -f, - slash_action_eq' := λ γ, by simpa [modular_form.subgroup_slash, - slash_action.neg_slash] using f.slash_action_eq' γ } ⟩ + slash_action_eq' := λ γ, by rw [slash_action.neg_slash, slash_action_eqn] } ⟩ @[simp] lemma coe_neg (f : slash_invariant_form Γ k) : ⇑(-f) = -f := rfl @[simp] lemma neg_apply (f : slash_invariant_form Γ k) (z : ℍ) : (-f) z = - (f z) := rfl From efed3cad43feb69219272b517de4aa4f21a81127 Mon Sep 17 00:00:00 2001 From: Thomas Browning Date: Fri, 5 May 2023 00:23:07 +0000 Subject: [PATCH 0956/1291] feat(field_theory/intermediate_field): Add lemma for the `field_range` of `val` (#18940) This PR adds a simple lemma for the `field_range` of `val`. --- src/field_theory/intermediate_field.lean | 3 +++ 1 file changed, 3 insertions(+) diff --git a/src/field_theory/intermediate_field.lean b/src/field_theory/intermediate_field.lean index 2306b210cb2a4..d23233382994f 100644 --- a/src/field_theory/intermediate_field.lean +++ b/src/field_theory/intermediate_field.lean @@ -349,6 +349,9 @@ S.to_subalgebra.val lemma range_val : S.val.range = S.to_subalgebra := S.to_subalgebra.range_val +@[simp] lemma field_range_val : S.val.field_range = S := +set_like.ext' subtype.range_val + lemma aeval_coe {R : Type*} [comm_ring R] [algebra R K] [algebra R L] [is_scalar_tower R K L] (x : S) (P : R[X]) : aeval (x : L) P = aeval x P := begin From 25580801f04aed44a0daf912d3760d0eaaf6d1bb Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Fri, 5 May 2023 01:40:46 +0000 Subject: [PATCH 0957/1291] chore(data/is_R_or_C/basic): rename `conj_mul_eq_norm_sq_left` to `conj_mul` (#18939) --- src/analysis/inner_product_space/basic.lean | 5 ++--- src/analysis/normed_space/extend.lean | 2 +- src/data/is_R_or_C/basic.lean | 10 ++-------- src/topology/continuous_function/ideals.lean | 2 +- 4 files changed, 6 insertions(+), 13 deletions(-) diff --git a/src/analysis/inner_product_space/basic.lean b/src/analysis/inner_product_space/basic.lean index 3ef38143a2648..361495b733589 100644 --- a/src/analysis/inner_product_space/basic.lean +++ b/src/analysis/inner_product_space/basic.lean @@ -259,7 +259,7 @@ begin rw [← @of_real_inj 𝕜, coe_norm_sq_eq_inner_self], simp only [inner_sub_sub_self, inner_smul_left, inner_smul_right, conj_of_real, mul_sub, ← coe_norm_sq_eq_inner_self x, ← coe_norm_sq_eq_inner_self y], - rw [← mul_assoc, mul_conj, is_R_or_C.conj_mul_eq_norm_sq_left, norm_sq_eq_def', + rw [← mul_assoc, mul_conj, is_R_or_C.conj_mul, norm_sq_eq_def', mul_left_comm, ← inner_conj_symm y, mul_conj, norm_sq_eq_def'], push_cast, ring @@ -327,8 +327,7 @@ def to_normed_space : normed_space 𝕜 F := { norm_smul_le := assume r x, begin rw [norm_eq_sqrt_inner, inner_smul_left, inner_smul_right, ←mul_assoc], - rw [is_R_or_C.conj_mul_eq_norm_sq_left, of_real_mul_re, sqrt_mul, ← coe_norm_sq_eq_inner_self, - of_real_re], + rw [is_R_or_C.conj_mul, of_real_mul_re, sqrt_mul, ← coe_norm_sq_eq_inner_self, of_real_re], { simp [sqrt_norm_sq_eq_norm, is_R_or_C.sqrt_norm_sq_eq_norm] }, { exact norm_sq_nonneg r } end } diff --git a/src/analysis/normed_space/extend.lean b/src/analysis/normed_space/extend.lean index bc01ce94c0d2b..ec74be9582adf 100644 --- a/src/analysis/normed_space/extend.lean +++ b/src/analysis/normed_space/extend.lean @@ -85,7 +85,7 @@ by simp only [extend_to_𝕜'_apply, map_sub, zero_mul, mul_zero, sub_zero] with lemma norm_extend_to_𝕜'_apply_sq (f : F →ₗ[ℝ] ℝ) (x : F) : ‖(f.extend_to_𝕜' x : 𝕜)‖ ^ 2 = f (conj (f.extend_to_𝕜' x : 𝕜) • x) := calc ‖(f.extend_to_𝕜' x : 𝕜)‖ ^ 2 = re (conj (f.extend_to_𝕜' x) * f.extend_to_𝕜' x : 𝕜) : - by rw [is_R_or_C.conj_mul_eq_norm_sq_left, norm_sq_eq_def', of_real_re] + by rw [is_R_or_C.conj_mul, norm_sq_eq_def', of_real_re] ... = f (conj (f.extend_to_𝕜' x : 𝕜) • x) : by rw [← smul_eq_mul, ← map_smul, extend_to_𝕜'_apply_re] diff --git a/src/data/is_R_or_C/basic.lean b/src/data/is_R_or_C/basic.lean index c4c1b2924aa12..2f8265c89a265 100644 --- a/src/data/is_R_or_C/basic.lean +++ b/src/data/is_R_or_C/basic.lean @@ -311,6 +311,8 @@ by simp only [map_add, add_zero, ext_iff, monoid_with_zero_hom.coe_mk, mul_neg, add_right_neg, zero_add, norm_sq, mul_comm, and_self, neg_neg, mul_zero, sub_eq_neg_add, neg_zero] with is_R_or_C_simps +lemma conj_mul (x : K) : conj x * x = ((norm_sq x) : K) := by rw [mul_comm, mul_conj] + theorem add_conj (z : K) : z + conj z = 2 * (re z) := by simp only [ext_iff, two_mul, map_add, add_zero, of_real_im, conj_im, of_real_re, eq_self_iff_true, add_right_neg, conj_re, and_self] @@ -590,14 +592,6 @@ lemma abs_sq_re_add_conj' (x : K) : (abs (conj x + x))^2 = (re (conj x + x))^2 : by simp only [sq, ←norm_sq_eq_abs, norm_sq, map_add, add_zero, monoid_with_zero_hom.coe_mk, add_left_neg, mul_zero] with is_R_or_C_simps -lemma conj_mul_eq_norm_sq_left (x : K) : conj x * x = ((norm_sq x) : K) := -begin - rw ext_iff, - refine ⟨by simp only [norm_sq, neg_mul, monoid_with_zero_hom.coe_mk, - sub_neg_eq_add, map_add, sub_zero, mul_zero] with is_R_or_C_simps, _⟩, - simp only [mul_comm, mul_neg, add_left_neg] with is_R_or_C_simps -end - /-! ### Cauchy sequences -/ theorem is_cau_seq_re (f : cau_seq K abs) : is_cau_seq abs' (λ n, re (f n)) := diff --git a/src/topology/continuous_function/ideals.lean b/src/topology/continuous_function/ideals.lean index 22779d09edb55..ece1cf99f5e39 100644 --- a/src/topology/continuous_function/ideals.lean +++ b/src/topology/continuous_function/ideals.lean @@ -260,7 +260,7 @@ begin ext, simp only [comp_apply, coe_mk, algebra_map_clm_coe, map_pow, coe_mul, coe_star, pi.mul_apply, pi.star_apply, star_def, continuous_map.coe_coe], - simpa only [norm_sq_eq_def', conj_mul_eq_norm_sq_left, of_real_pow], }, }, + simpa only [norm_sq_eq_def', is_R_or_C.conj_mul, of_real_pow], }, }, /- Get the function `g'` which is guaranteed to exist above. By the extreme value theorem and compactness of `t`, there is some `0 < c` such that `c ≤ g' x` for all `x ∈ t`. Then by `main_lemma_aux` there is some `g` for which `g * g'` is the desired function. -/ From 00f91228655eecdcd3ac97a7fd8dbcb139fe990a Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Fri, 5 May 2023 05:51:17 +0000 Subject: [PATCH 0958/1291] feat(number_theory/number_field/units): add is_unit_iff_norm (#18866) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR creates the file `number_theory/number_field/units.lean` and proves the result : ```lean lemma is_unit_iff_norm [number_field K] (x : 𝓞 K) : is_unit x ↔ abs (ring_of_integers.norm ℚ x : ℚ) = 1 ``` --- src/algebra/group_power/order.lean | 5 ++ src/algebra/ring/equiv.lean | 3 ++ src/field_theory/galois.lean | 10 ++++ .../is_alg_closed/algebraic_closure.lean | 1 + src/field_theory/is_alg_closed/basic.lean | 4 ++ src/field_theory/normal.lean | 3 ++ src/number_theory/number_field/norm.lean | 40 ++++++++++++++- src/number_theory/number_field/units.lean | 50 +++++++++++++++++++ 8 files changed, 114 insertions(+), 2 deletions(-) create mode 100644 src/number_theory/number_field/units.lean diff --git a/src/algebra/group_power/order.lean b/src/algebra/group_power/order.lean index 773cadffb22e5..d16d94f5abb1b 100644 --- a/src/algebra/group_power/order.lean +++ b/src/algebra/group_power/order.lean @@ -457,6 +457,11 @@ lemma pow_abs (a : R) (n : ℕ) : |a| ^ n = |a ^ n| := lemma abs_neg_one_pow (n : ℕ) : |(-1 : R) ^ n| = 1 := by rw [←pow_abs, abs_neg, abs_one, one_pow] +lemma abs_pow_eq_one (a : R) {n : ℕ} (h : 0 < n) : + |a ^ n| = 1 ↔ |a| = 1 := +by { convert pow_left_inj (abs_nonneg a) zero_le_one h, + exacts [(pow_abs _ _).symm, (one_pow _).symm] } + theorem pow_bit0_nonneg (a : R) (n : ℕ) : 0 ≤ a ^ bit0 n := by { rw pow_bit0, exact mul_self_nonneg _ } diff --git a/src/algebra/ring/equiv.lean b/src/algebra/ring/equiv.lean index 224d8eb3c521f..d68db88abdc04 100644 --- a/src/algebra/ring/equiv.lean +++ b/src/algebra/ring/equiv.lean @@ -392,6 +392,9 @@ variables [non_assoc_ring R] [non_assoc_ring S] (f : R ≃+* S) (x y : R) @[simp] lemma map_neg_one : f (-1) = -1 := f.map_one ▸ f.map_neg 1 +lemma map_eq_neg_one_iff {x : R} : f x = -1 ↔ x = -1 := +by rw [← neg_eq_iff_eq_neg, ← neg_eq_iff_eq_neg, ← map_neg, ring_equiv.map_eq_one_iff] + end ring section non_unital_semiring_hom diff --git a/src/field_theory/galois.lean b/src/field_theory/galois.lean index f818b32d57dba..5b4cc08749054 100644 --- a/src/field_theory/galois.lean +++ b/src/field_theory/galois.lean @@ -424,6 +424,16 @@ end is_galois end galois_equivalent_definitions +section normal_closure + +variables (k K F : Type*) [field k] [field K] [field F] [algebra k K] [algebra k F] + [algebra K F] [is_scalar_tower k K F] [is_galois k F] + +instance is_galois.normal_closure : is_galois k (normal_closure k K F) := +{ to_is_separable := is_separable_tower_bot_of_is_separable k _ F } + +end normal_closure + section is_alg_closure @[priority 100] diff --git a/src/field_theory/is_alg_closed/algebraic_closure.lean b/src/field_theory/is_alg_closed/algebraic_closure.lean index 6435bb6ed3c2a..18dc842228302 100644 --- a/src/field_theory/is_alg_closed/algebraic_closure.lean +++ b/src/field_theory/is_alg_closed/algebraic_closure.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau -/ import algebra.direct_limit +import algebra.char_p.algebra import field_theory.is_alg_closed.basic /-! diff --git a/src/field_theory/is_alg_closed/basic.lean b/src/field_theory/is_alg_closed/basic.lean index d04ad10dc000c..9765b18916db8 100644 --- a/src/field_theory/is_alg_closed/basic.lean +++ b/src/field_theory/is_alg_closed/basic.lean @@ -411,6 +411,10 @@ variables [algebra R S] [algebra R L] [is_scalar_tower R S L] variables [algebra K J] [algebra J L] [is_alg_closure J L] [algebra K L] [is_scalar_tower K J L] +/-- If `J` is an algebraic extension of `K` and `L` is an algebraic closure of `J`, then it is + also an algebraic closure of `K`. -/ +lemma of_algebraic (hKJ : algebra.is_algebraic K J) : is_alg_closure K L := +⟨is_alg_closure.alg_closed J, algebra.is_algebraic_trans hKJ is_alg_closure.algebraic⟩ /-- A (random) isomorphism between an algebraic closure of `R` and an algebraic closure of an algebraic extension of `R` -/ diff --git a/src/field_theory/normal.lean b/src/field_theory/normal.lean index eac698418193c..107129a8dbb4f 100644 --- a/src/field_theory/normal.lean +++ b/src/field_theory/normal.lean @@ -435,6 +435,9 @@ begin apply intermediate_field.finite_dimensional_supr_of_finite, end +instance is_scalar_tower : is_scalar_tower F (normal_closure F K L) L := +is_scalar_tower.subalgebra' F L L _ + end normal_closure end normal_closure diff --git a/src/number_theory/number_field/norm.lean b/src/number_theory/number_field/norm.lean index 35428dcf00f67..8476cfc5dff86 100644 --- a/src/number_theory/number_field/norm.lean +++ b/src/number_theory/number_field/norm.lean @@ -22,7 +22,7 @@ rings of integers. open_locale number_field big_operators -open finset number_field algebra +open finset number_field algebra finite_dimensional namespace ring_of_integers @@ -37,7 +37,15 @@ local attribute [instance] number_field.ring_of_integers_algebra lemma coe_algebra_map_norm [is_separable K L] (x : 𝓞 L) : (algebra_map (𝓞 K) (𝓞 L) (norm K x) : L) = algebra_map K L (algebra.norm K (x : L)) := rfl -lemma is_unit_norm [is_galois K L] {x : 𝓞 L} : +lemma coe_norm_algebra_map [is_separable K L] (x : 𝓞 K) : + (norm K (algebra_map (𝓞 K) (𝓞 L) x) : K) = algebra.norm K (algebra_map K L x) := rfl + +lemma norm_algebra_map [is_separable K L] (x : 𝓞 K) : + norm K (algebra_map (𝓞 K) (𝓞 L) x) = x ^ finrank K L := +by rw [← subtype.coe_inj, ring_of_integers.coe_norm_algebra_map, algebra.norm_algebra_map, + subsemiring_class.coe_pow] + +lemma is_unit_norm_of_is_galois [is_galois K L] {x : 𝓞 L} : is_unit (norm K x) ↔ is_unit x := begin classical, @@ -68,4 +76,32 @@ begin simp [← finset.mul_prod_erase _ _ (mem_univ alg_equiv.refl)] end +variables (F : Type*) [field F] [algebra K F] [is_separable K F] [finite_dimensional K F] + +lemma norm_norm [is_separable K L] [algebra F L] [is_separable F L] [finite_dimensional F L] + [is_scalar_tower K F L] (x : 𝓞 L) : norm K (norm F x) = norm K x := +by rw [← subtype.coe_inj, norm_apply_coe, norm_apply_coe, norm_apply_coe, algebra.norm_norm] + +variable {F} + +lemma is_unit_norm [char_zero K] {x : 𝓞 F} : + is_unit (norm K x) ↔ is_unit x := +begin + letI : algebra K (algebraic_closure K) := algebraic_closure.algebra K, + let L := normal_closure K F (algebraic_closure F), + haveI : finite_dimensional F L := finite_dimensional.right K F L, + haveI : is_alg_closure K (algebraic_closure F) := + is_alg_closure.of_algebraic K F (algebraic_closure F) (algebra.is_algebraic_of_finite K F), + haveI : is_galois F L := is_galois.tower_top_of_is_galois K F L, + calc + is_unit (norm K x) ↔ is_unit ((norm K) x ^ finrank F L) : + (is_unit_pow_iff (pos_iff_ne_zero.mp finrank_pos)).symm + ... ↔ is_unit (norm K (algebra_map (𝓞 F) (𝓞 L) x)) : + by rw [← norm_norm K F (algebra_map (𝓞 F) (𝓞 L) x), norm_algebra_map F _, map_pow] + ... ↔ is_unit (algebra_map (𝓞 F) (𝓞 L) x) : is_unit_norm_of_is_galois K + ... ↔ is_unit (norm F (algebra_map (𝓞 F) (𝓞 L) x)) : (is_unit_norm_of_is_galois F).symm + ... ↔ is_unit (x ^ finrank F L) : (congr_arg is_unit (norm_algebra_map F _)).to_iff + ... ↔ is_unit x : is_unit_pow_iff (pos_iff_ne_zero.mp finrank_pos), +end + end ring_of_integers diff --git a/src/number_theory/number_field/units.lean b/src/number_theory/number_field/units.lean new file mode 100644 index 0000000000000..106abb57e6437 --- /dev/null +++ b/src/number_theory/number_field/units.lean @@ -0,0 +1,50 @@ +/- +Copyright (c) 2023 Xavier Roblot. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Xavier Roblot +-/ +import number_theory.number_field.norm + +/-! +# Units of a number field +We prove results about the group `(𝓞 K)ˣ` of units of the ring of integers `𝓞 K` of a number +field `K`. + +## Main results +* `number_field.is_unit_iff_norm`: an algebraic integer `x : 𝓞 K` is a unit if and only if +`|norm ℚ x| = 1` + +## Tags +number field, units + -/ + +open_locale number_field + +noncomputable theory + +open number_field units + +section rat + +lemma rat.ring_of_integers.is_unit_iff {x : 𝓞 ℚ} : + is_unit x ↔ ((x : ℚ) = 1) ∨ ((x : ℚ) = -1) := +by simp_rw [(is_unit_map_iff (rat.ring_of_integers_equiv : 𝓞 ℚ →+* ℤ) x).symm, int.is_unit_iff, + ring_equiv.coe_to_ring_hom, ring_equiv.map_eq_one_iff, ring_equiv.map_eq_neg_one_iff, + ← subtype.coe_injective.eq_iff, add_subgroup_class.coe_neg, algebra_map.coe_one] + +end rat + +variables (K : Type*) [field K] + +section is_unit + +local attribute [instance] number_field.ring_of_integers_algebra + +variable {K} + +lemma is_unit_iff_norm [number_field K] (x : 𝓞 K) : + is_unit x ↔ |(ring_of_integers.norm ℚ x : ℚ)| = 1 := +by { convert (ring_of_integers.is_unit_norm ℚ).symm, + rw [← abs_one, abs_eq_abs, ← rat.ring_of_integers.is_unit_iff], } + +end is_unit From 3905fa80e62c0898131285baab35559fbc4e5cda Mon Sep 17 00:00:00 2001 From: Felix-Weilacher Date: Fri, 5 May 2023 08:50:18 +0000 Subject: [PATCH 0959/1291] doc(measure_theory): Fix some names and docstrings (#18910) Co-authored-by: Felix-Weilacher <112423742+Felix-Weilacher@users.noreply.github.com> --- src/measure_theory/constructions/polish.lean | 9 ++++++--- src/measure_theory/measurable_space.lean | 2 +- src/topology/perfect.lean | 4 +--- 3 files changed, 8 insertions(+), 7 deletions(-) diff --git a/src/measure_theory/constructions/polish.lean b/src/measure_theory/constructions/polish.lean index 685cab7f6fd2e..546b85d2501a0 100644 --- a/src/measure_theory/constructions/polish.lean +++ b/src/measure_theory/constructions/polish.lean @@ -47,8 +47,10 @@ analytic sets. We use this to prove several versions of the Borel isomorphism theorem. -* `measurable_equiv_of_not_countable` : Any two uncountable Polish spaces are Borel isomorphic. -* `equiv.measurable_equiv` : Any two Polish spaces of the same cardinality are Borel. isomorphic. +* `polish_space.measurable_equiv_of_not_countable` : Any two uncountable Polish spaces + are Borel isomorphic. +* `polish_space.equiv.measurable_equiv` : Any two Polish spaces of the same cardinality + are Borel isomorphic. -/ open set function polish_space pi_nat topological_space metric filter @@ -760,7 +762,8 @@ begin apply nonempty.some, obtain ⟨f, -, fcts, finj⟩ := is_closed_univ.exists_nat_bool_injection_of_not_countable (by rwa [← countable_coe_iff, (equiv.set.univ _).countable_iff]), - obtain ⟨g, gmeas, ginj⟩ := measurable_space.measurable_injection_cantor_of_countably_generated α, + obtain ⟨g, gmeas, ginj⟩ := + measurable_space.measurable_injection_nat_bool_of_countably_generated α, exact ⟨borel_schroeder_bernstein gmeas ginj fcts.measurable finj⟩, end diff --git a/src/measure_theory/measurable_space.lean b/src/measure_theory/measurable_space.lean index a51e01907c94b..60e479ef9cf97 100644 --- a/src/measure_theory/measurable_space.lean +++ b/src/measure_theory/measurable_space.lean @@ -1392,7 +1392,7 @@ open_locale classical /-- If a measurable space is countably generated, it admits a measurable injection into the Cantor space `ℕ → bool` (equipped with the product sigma algebra). -/ -theorem measurable_injection_cantor_of_countably_generated +theorem measurable_injection_nat_bool_of_countably_generated [measurable_space α] [h : countably_generated α] [measurable_singleton_class α] : ∃ f : α → (ℕ → bool), measurable f ∧ function.injective f := begin diff --git a/src/topology/perfect.lean b/src/topology/perfect.lean index 1321cd68d34d8..65a5f51d1df9c 100644 --- a/src/topology/perfect.lean +++ b/src/topology/perfect.lean @@ -19,8 +19,6 @@ including a version of the Cantor-Bendixson Theorem. * `perfect C`: A set `C` is perfect, meaning it is closed and every point of it is an accumulation point of itself. -* `set.scheme β α`: A `β`-scheme on `α`, a collection of subsets of `α` indexed by `list β`. - Used to construct maps `(β → ℕ) → α` as limiting objects. ## Main Statements @@ -30,7 +28,7 @@ including a version of the Cantor-Bendixson Theorem. * `exists_countable_union_perfect_of_is_closed`: One version of the **Cantor-Bendixson Theorem**: A closed set in a second countable space can be written as the union of a countable set and a perfect set. -* `exists_nat_bool_injection_of_perfect_nonempty`: A perfect nonempty set in a complete metric space +* `perfect.exists_nat_bool_injection`: A perfect nonempty set in a complete metric space admits an embedding from the Cantor space. ## Implementation Notes From fa4a805d16a9cd9c96e0f8edeb57dc5a07af1a19 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 6 May 2023 08:38:22 +0000 Subject: [PATCH 0960/1291] =?UTF-8?q?feat(order/category):=20Forgetful=20f?= =?UTF-8?q?unctor=20`NonemptyFinLinOrd=20=E2=A5=A4=20FinPartOrd`=20(#18948?= =?UTF-8?q?)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Also fix a wrong docstring --- src/order/category/NonemptyFinLinOrd.lean | 15 ++++++++++++++- 1 file changed, 14 insertions(+), 1 deletion(-) diff --git a/src/order/category/NonemptyFinLinOrd.lean b/src/order/category/NonemptyFinLinOrd.lean index db8e1d206bc0d..73bddd81f550e 100644 --- a/src/order/category/NonemptyFinLinOrd.lean +++ b/src/order/category/NonemptyFinLinOrd.lean @@ -5,6 +5,7 @@ Authors: Johan Commelin -/ import data.fintype.order import data.set.finite +import order.category.FinPartOrd import order.category.LinOrd import category_theory.limits.shapes.images import category_theory.limits.shapes.regular_mono @@ -17,6 +18,9 @@ import category_theory.limits.shapes.regular_mono This defines `NonemptyFinLinOrd`, the category of nonempty finite linear orders with monotone maps. This is the index category for simplicial objects. + +Note: `NonemptyFinLinOrd` is *not* a subcategory of `FinBddDistLat` because its morphisms do not +preserve `⊥` and `⊤`. -/ universes u v @@ -68,6 +72,9 @@ instance (α : NonemptyFinLinOrd) : nonempty_fin_lin_ord α := α.str instance has_forget_to_LinOrd : has_forget₂ NonemptyFinLinOrd LinOrd := bundled_hom.forget₂ _ _ +instance has_forget_to_FinPartOrd : has_forget₂ NonemptyFinLinOrd FinPartOrd := +{ forget₂ := { obj := λ X, FinPartOrd.of X, map := λ X Y, id } } + /-- Constructs an equivalence between nonempty finite linear orders from an order isomorphism between them. -/ @[simps] def iso.mk {α β : NonemptyFinLinOrd.{u}} (e : α ≃o β) : α ≅ β := @@ -80,7 +87,7 @@ between them. -/ @[simps] def dual : NonemptyFinLinOrd ⥤ NonemptyFinLinOrd := { obj := λ X, of Xᵒᵈ, map := λ X Y, order_hom.dual } -/-- The equivalence between `FinPartOrd` and itself induced by `order_dual` both ways. -/ +/-- The equivalence between `NonemptyFinLinOrd` and itself induced by `order_dual` both ways. -/ @[simps functor inverse] def dual_equiv : NonemptyFinLinOrd ≌ NonemptyFinLinOrd := equivalence.mk dual dual (nat_iso.of_components (λ X, iso.mk $ order_iso.dual_dual X) $ λ X Y f, rfl) @@ -182,3 +189,9 @@ end NonemptyFinLinOrd lemma NonemptyFinLinOrd_dual_comp_forget_to_LinOrd : NonemptyFinLinOrd.dual ⋙ forget₂ NonemptyFinLinOrd LinOrd = forget₂ NonemptyFinLinOrd LinOrd ⋙ LinOrd.dual := rfl + +/-- The forgetful functor `NonemptyFinLinOrd ⥤ FinPartOrd` and `order_dual` commute. -/ +def NonemptyFinLinOrd_dual_comp_forget_to_FinPartOrd : + NonemptyFinLinOrd.dual ⋙ forget₂ NonemptyFinLinOrd FinPartOrd ≅ + forget₂ NonemptyFinLinOrd FinPartOrd ⋙ FinPartOrd.dual := +{ hom := { app := λ X, order_hom.id }, inv := { app := λ X, order_hom.id } } From db07e6feaf211fe704c1e79ba5f480fd6d218523 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Sat, 6 May 2023 10:26:45 +0000 Subject: [PATCH 0961/1291] chore(algebra/order/lattice_group): remove redundant instance (#18951) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This instance was redundant. It can be synthesised at the declaration site via: ``` example (α : Type u) [linear_ordered_comm_group α] : covariant_class α α (*) (≤) := by show_term {apply_instance} -- ordered_comm_group.to_covariant_class_left_le α ``` Moreover, it is implicated in a timeout in the reenableeta branches, so I want it gone! This PR is just a verification that mathlib still compiles. I don't care if it is merged. Co-authored-by: Scott Morrison --- src/algebra/order/lattice_group.lean | 6 ------ 1 file changed, 6 deletions(-) diff --git a/src/algebra/order/lattice_group.lean b/src/algebra/order/lattice_group.lean index b1cd7376dd9a7..0dbc871a56fc8 100644 --- a/src/algebra/order/lattice_group.lean +++ b/src/algebra/order/lattice_group.lean @@ -62,12 +62,6 @@ lattice, ordered, group universe u --- A linearly ordered additive commutative group is a lattice ordered commutative group -@[priority 100, to_additive] -- see Note [lower instance priority] -instance linear_ordered_comm_group.to_covariant_class (α : Type u) - [linear_ordered_comm_group α] : covariant_class α α (*) (≤) := -{ elim := λ a b c bc, linear_ordered_comm_group.mul_le_mul_left _ _ bc a } - variables {α : Type u} [lattice α] [comm_group α] -- Special case of Bourbaki A.VI.9 (1) From 4c8f86bdbe06f92960d0eb314da8ca64a9d33bca Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sat, 6 May 2023 13:10:53 +0000 Subject: [PATCH 0962/1291] feat(algebra/order/to_interval_mod): symmetric variants of lemmas (#18942) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit These lemmas are about expressions in the second instead of third arguments of the `I{co,oc}_{mod,div}` functions. Some existing lemmas clashed with these new lemmas; they have been renamed as follows: * `to_Ico_div_sub'` → `to_Ico_div_sub_eq_to_Ico_div_add` * `to_Ioc_div_sub'` → `to_Ioc_div_sub_eq_to_Ioc_div_add` * `to_Ico_div_add_right'` → `to_Ico_div_sub_eq_to_Ico_div_add'` (and reversed) * `to_Ioc_div_add_right'` → `to_Ioc_div_sub_eq_to_Ioc_div_add'` (and reversed) * `to_Ico_mod_sub'` → `to_Ico_mod_sub_eq_sub` * `to_Ioc_mod_sub'` → `to_Ioc_mod_sub_eq_sub` * `to_Ico_mod_add_right'` → `to_Ico_mod_add_right_eq_add` * `to_Ioc_mod_add_right'` → `to_Ioc_mod_add_right_eq_add` The statement of `to_Ioc_div_zsmul_add` is commuted to be consistent with the lemmas around it; presumably it was a copy-and-paste error. --- src/algebra/order/to_interval_mod.lean | 139 ++++++++++++++++++++++--- 1 file changed, 123 insertions(+), 16 deletions(-) diff --git a/src/algebra/order/to_interval_mod.lean b/src/algebra/order/to_interval_mod.lean index 4943b56adce75..c423f022f6fc2 100644 --- a/src/algebra/order/to_interval_mod.lean +++ b/src/algebra/order/to_interval_mod.lean @@ -181,69 +181,120 @@ by { rw [to_Ioc_mod_eq_iff hp, set.right_mem_Ioc], exact ⟨lt_add_of_pos_right to_Ico_div_eq_of_sub_zsmul_mem_Ico hp $ by simpa only [add_smul, add_sub_add_right_eq_sub] using sub_to_Ico_div_zsmul_mem_Ico hp a b +@[simp] lemma to_Ico_div_add_zsmul' (a b: α) (m : ℤ) : + to_Ico_div hp (a + m • p) b = to_Ico_div hp a b - m := +begin + refine to_Ico_div_eq_of_sub_zsmul_mem_Ico _ _, + rw [sub_smul, ←sub_add, add_right_comm], + simpa using sub_to_Ico_div_zsmul_mem_Ico hp a b, +end + @[simp] lemma to_Ioc_div_add_zsmul (a b : α) (m : ℤ) : to_Ioc_div hp a (b + m • p) = to_Ioc_div hp a b + m := to_Ioc_div_eq_of_sub_zsmul_mem_Ioc hp $ by simpa only [add_smul, add_sub_add_right_eq_sub] using sub_to_Ioc_div_zsmul_mem_Ioc hp a b +@[simp] lemma to_Ioc_div_add_zsmul' (a b : α) (m : ℤ) : + to_Ioc_div hp (a + m • p) b = to_Ioc_div hp a b - m := +begin + refine to_Ioc_div_eq_of_sub_zsmul_mem_Ioc _ _, + rw [sub_smul, ←sub_add, add_right_comm], + simpa using sub_to_Ioc_div_zsmul_mem_Ioc hp a b, +end + @[simp] lemma to_Ico_div_zsmul_add (a b : α) (m : ℤ) : to_Ico_div hp a (m • p + b) = m + to_Ico_div hp a b := by rw [add_comm, to_Ico_div_add_zsmul, add_comm] +/-! Note we omit `to_Ico_div_zsmul_add'` as `-m + to_Ico_div hp a b` is not very convenient. -/ + @[simp] lemma to_Ioc_div_zsmul_add (a b : α) (m : ℤ) : - to_Ioc_div hp a (m • p + b) = to_Ioc_div hp a b + m := + to_Ioc_div hp a (m • p + b) = m + to_Ioc_div hp a b := by rw [add_comm, to_Ioc_div_add_zsmul, add_comm] +/-! Note we omit `to_Ioc_div_zsmul_add'` as `-m + to_Ioc_div hp a b` is not very convenient. -/ + @[simp] lemma to_Ico_div_sub_zsmul (a b : α) (m : ℤ) : to_Ico_div hp a (b - m • p) = to_Ico_div hp a b - m := by rw [sub_eq_add_neg, ←neg_smul, to_Ico_div_add_zsmul, sub_eq_add_neg] +@[simp] lemma to_Ico_div_sub_zsmul' (a b : α) (m : ℤ) : + to_Ico_div hp (a - m • p) b = to_Ico_div hp a b + m := +by rw [sub_eq_add_neg, ←neg_smul, to_Ico_div_add_zsmul', sub_neg_eq_add] + @[simp] lemma to_Ioc_div_sub_zsmul (a b : α) (m : ℤ) : to_Ioc_div hp a (b - m • p) = to_Ioc_div hp a b - m := by rw [sub_eq_add_neg, ←neg_smul, to_Ioc_div_add_zsmul, sub_eq_add_neg] +@[simp] lemma to_Ioc_div_sub_zsmul' (a b : α) (m : ℤ) : + to_Ioc_div hp (a - m • p) b = to_Ioc_div hp a b + m := +by rw [sub_eq_add_neg, ←neg_smul, to_Ioc_div_add_zsmul', sub_neg_eq_add] + @[simp] lemma to_Ico_div_add_right (a b : α) : to_Ico_div hp a (b + p) = to_Ico_div hp a b + 1 := by simpa only [one_zsmul] using to_Ico_div_add_zsmul hp a b 1 +@[simp] lemma to_Ico_div_add_right' (a b : α) : to_Ico_div hp (a + p) b = to_Ico_div hp a b - 1 := +by simpa only [one_zsmul] using to_Ico_div_add_zsmul' hp a b 1 + @[simp] lemma to_Ioc_div_add_right (a b : α) : to_Ioc_div hp a (b + p) = to_Ioc_div hp a b + 1 := by simpa only [one_zsmul] using to_Ioc_div_add_zsmul hp a b 1 +@[simp] lemma to_Ioc_div_add_right' (a b : α) : to_Ioc_div hp (a + p) b = to_Ioc_div hp a b - 1 := +by simpa only [one_zsmul] using to_Ioc_div_add_zsmul' hp a b 1 + @[simp] lemma to_Ico_div_add_left (a b : α) : to_Ico_div hp a (p + b) = to_Ico_div hp a b + 1 := by rw [add_comm, to_Ico_div_add_right] +@[simp] lemma to_Ico_div_add_left' (a b : α) : to_Ico_div hp (p + a) b = to_Ico_div hp a b - 1 := +by rw [add_comm, to_Ico_div_add_right'] + @[simp] lemma to_Ioc_div_add_left (a b : α) : to_Ioc_div hp a (p + b) = to_Ioc_div hp a b + 1 := by rw [add_comm, to_Ioc_div_add_right] +@[simp] lemma to_Ioc_div_add_left' (a b : α) : to_Ioc_div hp (p + a) b = to_Ioc_div hp a b - 1 := +by rw [add_comm, to_Ioc_div_add_right'] + @[simp] lemma to_Ico_div_sub (a b : α) : to_Ico_div hp a (b - p) = to_Ico_div hp a b - 1 := by simpa only [one_zsmul] using to_Ico_div_sub_zsmul hp a b 1 +@[simp] lemma to_Ico_div_sub' (a b : α) : to_Ico_div hp (a - p) b = to_Ico_div hp a b + 1 := +by simpa only [one_zsmul] using to_Ico_div_sub_zsmul' hp a b 1 + @[simp] lemma to_Ioc_div_sub (a b : α) : to_Ioc_div hp a (b - p) = to_Ioc_div hp a b - 1 := by simpa only [one_zsmul] using to_Ioc_div_sub_zsmul hp a b 1 -lemma to_Ico_div_sub' (a b c : α) : to_Ico_div hp a (b - c) = to_Ico_div hp (a + c) b := +@[simp] lemma to_Ioc_div_sub' (a b : α) : to_Ioc_div hp (a - p) b = to_Ioc_div hp a b + 1 := +by simpa only [one_zsmul] using to_Ioc_div_sub_zsmul' hp a b 1 + +lemma to_Ico_div_sub_eq_to_Ico_div_add (a b c : α) : + to_Ico_div hp a (b - c) = to_Ico_div hp (a + c) b := begin apply to_Ico_div_eq_of_sub_zsmul_mem_Ico, rw [←sub_right_comm, set.sub_mem_Ico_iff_left, add_right_comm], exact sub_to_Ico_div_zsmul_mem_Ico hp (a + c) b, end -lemma to_Ioc_div_sub' (a b c : α) : to_Ioc_div hp a (b - c) = to_Ioc_div hp (a + c) b := +lemma to_Ioc_div_sub_eq_to_Ioc_div_add (a b c : α) : + to_Ioc_div hp a (b - c) = to_Ioc_div hp (a + c) b := begin apply to_Ioc_div_eq_of_sub_zsmul_mem_Ioc, rw [←sub_right_comm, set.sub_mem_Ioc_iff_left, add_right_comm], exact sub_to_Ioc_div_zsmul_mem_Ioc hp (a + c) b, end -lemma to_Ico_div_add_right' (a b c : α) : to_Ico_div hp a (b + c) = to_Ico_div hp (a - c) b := -by rw [←sub_neg_eq_add, to_Ico_div_sub', sub_eq_add_neg] +lemma to_Ico_div_sub_eq_to_Ico_div_add' (a b c : α) : + to_Ico_div hp (a - c) b = to_Ico_div hp a (b + c) := +by rw [←sub_neg_eq_add, to_Ico_div_sub_eq_to_Ico_div_add, sub_eq_add_neg] -lemma to_Ioc_div_add_right' (a b c : α) : to_Ioc_div hp a (b + c) = to_Ioc_div hp (a - c) b := -by rw [←sub_neg_eq_add, to_Ioc_div_sub', sub_eq_add_neg] +lemma to_Ioc_div_sub_eq_to_Ioc_div_add' (a b c : α) : + to_Ioc_div hp (a - c) b = to_Ioc_div hp a (b + c) := +by rw [←sub_neg_eq_add, to_Ioc_div_sub_eq_to_Ioc_div_add, sub_eq_add_neg] lemma to_Ico_div_neg (a b : α) : to_Ico_div hp a (-b) = -(to_Ioc_div hp (-a) b + 1) := begin suffices : to_Ico_div hp a (-b) = -(to_Ioc_div hp (-(a + p)) b), - { rwa [neg_add, ←sub_eq_add_neg, ←to_Ioc_div_add_right', to_Ioc_div_add_right] at this }, + { rwa [neg_add, ←sub_eq_add_neg, to_Ioc_div_sub_eq_to_Ioc_div_add', + to_Ioc_div_add_right] at this }, rw [← neg_eq_iff_eq_neg, eq_comm], apply to_Ioc_div_eq_of_sub_zsmul_mem_Ioc, obtain ⟨hc, ho⟩ := sub_to_Ico_div_zsmul_mem_Ico hp a (-b), @@ -253,69 +304,125 @@ begin rw [neg_add, neg_add_cancel_right], end +lemma to_Ico_div_neg' (a b : α) : to_Ico_div hp (-a) b = -(to_Ioc_div hp a (-b) + 1) := +by simpa only [neg_neg] using to_Ico_div_neg hp (-a) (-b) + lemma to_Ioc_div_neg (a b : α) : to_Ioc_div hp a (-b) = -(to_Ico_div hp (-a) b + 1) := by rw [←neg_neg b, to_Ico_div_neg, neg_neg, neg_neg, neg_add', neg_neg, add_sub_cancel] +lemma to_Ioc_div_neg' (a b : α) : to_Ioc_div hp (-a) b = -(to_Ico_div hp a (-b) + 1) := +by simpa only [neg_neg] using to_Ioc_div_neg hp (-a) (-b) + @[simp] lemma to_Ico_mod_add_zsmul (a b : α) (m : ℤ) : to_Ico_mod hp a (b + m • p) = to_Ico_mod hp a b := by { rw [to_Ico_mod, to_Ico_div_add_zsmul, to_Ico_mod, add_smul], abel } +@[simp] lemma to_Ico_mod_add_zsmul' (a b : α) (m : ℤ) : + to_Ico_mod hp (a + m • p) b = to_Ico_mod hp a b + m • p := +by simp only [to_Ico_mod, to_Ico_div_add_zsmul', sub_smul, sub_add] + @[simp] lemma to_Ioc_mod_add_zsmul (a b : α) (m : ℤ) : to_Ioc_mod hp a (b + m • p) = to_Ioc_mod hp a b := by { rw [to_Ioc_mod, to_Ioc_div_add_zsmul, to_Ioc_mod, add_smul], abel } +@[simp] lemma to_Ioc_mod_add_zsmul' (a b : α) (m : ℤ) : + to_Ioc_mod hp (a + m • p) b = to_Ioc_mod hp a b + m • p := +by simp only [to_Ioc_mod, to_Ioc_div_add_zsmul', sub_smul, sub_add] + @[simp] lemma to_Ico_mod_zsmul_add (a b : α) (m : ℤ) : to_Ico_mod hp a (m • p + b) = to_Ico_mod hp a b := by rw [add_comm, to_Ico_mod_add_zsmul] +@[simp] lemma to_Ico_mod_zsmul_add' (a b : α) (m : ℤ) : + to_Ico_mod hp (m • p + a) b = m • p + to_Ico_mod hp a b := +by rw [add_comm, to_Ico_mod_add_zsmul', add_comm] + @[simp] lemma to_Ioc_mod_zsmul_add (a b : α) (m : ℤ) : to_Ioc_mod hp a (m • p + b) = to_Ioc_mod hp a b := by rw [add_comm, to_Ioc_mod_add_zsmul] +@[simp] lemma to_Ioc_mod_zsmul_add' (a b : α) (m : ℤ) : + to_Ioc_mod hp (m • p + a) b = m • p + to_Ioc_mod hp a b := +by rw [add_comm, to_Ioc_mod_add_zsmul', add_comm] + @[simp] lemma to_Ico_mod_sub_zsmul (a b : α) (m : ℤ) : to_Ico_mod hp a (b - m • p) = to_Ico_mod hp a b := by rw [sub_eq_add_neg, ←neg_smul, to_Ico_mod_add_zsmul] +@[simp] lemma to_Ico_mod_sub_zsmul' (a b : α) (m : ℤ) : + to_Ico_mod hp (a - m • p) b = to_Ico_mod hp a b - m • p := +by simp_rw [sub_eq_add_neg, ←neg_smul, to_Ico_mod_add_zsmul'] + @[simp] lemma to_Ioc_mod_sub_zsmul (a b : α) (m : ℤ) : to_Ioc_mod hp a (b - m • p) = to_Ioc_mod hp a b := by rw [sub_eq_add_neg, ←neg_smul, to_Ioc_mod_add_zsmul] +@[simp] lemma to_Ioc_mod_sub_zsmul' (a b : α) (m : ℤ) : + to_Ioc_mod hp (a - m • p) b = to_Ioc_mod hp a b - m • p := +by simp_rw [sub_eq_add_neg, ←neg_smul, to_Ioc_mod_add_zsmul'] + @[simp] lemma to_Ico_mod_add_right (a b : α) : to_Ico_mod hp a (b + p) = to_Ico_mod hp a b := by simpa only [one_zsmul] using to_Ico_mod_add_zsmul hp a b 1 +@[simp] lemma to_Ico_mod_add_right' (a b : α) : to_Ico_mod hp (a + p) b = to_Ico_mod hp a b + p := +by simpa only [one_zsmul] using to_Ico_mod_add_zsmul' hp a b 1 + @[simp] lemma to_Ioc_mod_add_right (a b : α) : to_Ioc_mod hp a (b + p) = to_Ioc_mod hp a b := by simpa only [one_zsmul] using to_Ioc_mod_add_zsmul hp a b 1 +@[simp] lemma to_Ioc_mod_add_right' (a b : α) : to_Ioc_mod hp (a + p) b = to_Ioc_mod hp a b + p := +by simpa only [one_zsmul] using to_Ioc_mod_add_zsmul' hp a b 1 + @[simp] lemma to_Ico_mod_add_left (a b : α) : to_Ico_mod hp a (p + b) = to_Ico_mod hp a b := by rw [add_comm, to_Ico_mod_add_right] +@[simp] lemma to_Ico_mod_add_left' (a b : α) : to_Ico_mod hp (p + a) b = p + to_Ico_mod hp a b := +by rw [add_comm, to_Ico_mod_add_right', add_comm] + @[simp] lemma to_Ioc_mod_add_left (a b : α) : to_Ioc_mod hp a (p + b) = to_Ioc_mod hp a b := by rw [add_comm, to_Ioc_mod_add_right] +@[simp] lemma to_Ioc_mod_add_left' (a b : α) : to_Ioc_mod hp (p + a) b = p + to_Ioc_mod hp a b := +by rw [add_comm, to_Ioc_mod_add_right', add_comm] + @[simp] lemma to_Ico_mod_sub (a b : α) : to_Ico_mod hp a (b - p) = to_Ico_mod hp a b := by simpa only [one_zsmul] using to_Ico_mod_sub_zsmul hp a b 1 +@[simp] lemma to_Ico_mod_sub' (a b : α) : to_Ico_mod hp (a - p) b = to_Ico_mod hp a b - p := +by simpa only [one_zsmul] using to_Ico_mod_sub_zsmul' hp a b 1 + @[simp] lemma to_Ioc_mod_sub (a b : α) : to_Ioc_mod hp a (b - p) = to_Ioc_mod hp a b := by simpa only [one_zsmul] using to_Ioc_mod_sub_zsmul hp a b 1 -lemma to_Ico_mod_sub' (a b c : α) : to_Ico_mod hp a (b - c) = to_Ico_mod hp (a + c) b - c := -by simp_rw [to_Ico_mod, to_Ico_div_sub', sub_right_comm] +@[simp] lemma to_Ioc_mod_sub' (a b : α) : to_Ioc_mod hp (a - p) b = to_Ioc_mod hp a b - p := +by simpa only [one_zsmul] using to_Ioc_mod_sub_zsmul' hp a b 1 + +lemma to_Ico_mod_sub_eq_sub (a b c : α) : to_Ico_mod hp a (b - c) = to_Ico_mod hp (a + c) b - c := +by simp_rw [to_Ico_mod, to_Ico_div_sub_eq_to_Ico_div_add, sub_right_comm] -lemma to_Ioc_mod_sub' (a b c : α) : to_Ioc_mod hp a (b - c) = to_Ioc_mod hp (a + c) b - c := -by simp_rw [to_Ioc_mod, to_Ioc_div_sub', sub_right_comm] +lemma to_Ioc_mod_sub_eq_sub (a b c : α) : to_Ioc_mod hp a (b - c) = to_Ioc_mod hp (a + c) b - c := +by simp_rw [to_Ioc_mod, to_Ioc_div_sub_eq_to_Ioc_div_add, sub_right_comm] -lemma to_Ico_mod_add_right' (a b c : α) : to_Ico_mod hp a (b + c) = to_Ico_mod hp (a - c) b + c := -by simp_rw [to_Ico_mod, to_Ico_div_add_right', sub_add_eq_add_sub] +lemma to_Ico_mod_add_right_eq_add (a b c : α) : + to_Ico_mod hp a (b + c) = to_Ico_mod hp (a - c) b + c := +by simp_rw [to_Ico_mod, to_Ico_div_sub_eq_to_Ico_div_add', sub_add_eq_add_sub] -lemma to_Ioc_mod_add_right' (a b c : α) : to_Ioc_mod hp a (b + c) = to_Ioc_mod hp (a - c) b + c := -by simp_rw [to_Ioc_mod, to_Ioc_div_add_right', sub_add_eq_add_sub] +lemma to_Ioc_mod_add_right_eq_add (a b c : α) : + to_Ioc_mod hp a (b + c) = to_Ioc_mod hp (a - c) b + c := +by simp_rw [to_Ioc_mod, to_Ioc_div_sub_eq_to_Ioc_div_add', sub_add_eq_add_sub] lemma to_Ico_mod_neg (a b : α) : to_Ico_mod hp a (-b) = p - to_Ioc_mod hp (-a) b := by { simp_rw [to_Ico_mod, to_Ioc_mod, to_Ico_div_neg, neg_smul, add_smul], abel } +lemma to_Ico_mod_neg' (a b : α) : to_Ico_mod hp (-a) b = p - to_Ioc_mod hp a (-b) := +by simpa only [neg_neg] using to_Ico_mod_neg hp (-a) (-b) + lemma to_Ioc_mod_neg (a b : α) : to_Ioc_mod hp a (-b) = p - to_Ico_mod hp (-a) b := by { simp_rw [to_Ioc_mod, to_Ico_mod, to_Ioc_div_neg, neg_smul, add_smul], abel } +lemma to_Ioc_mod_neg' (a b : α) : to_Ioc_mod hp (-a) b = p - to_Ico_mod hp a (-b) := +by simpa only [neg_neg] using to_Ioc_mod_neg hp (-a) (-b) + lemma to_Ico_mod_eq_to_Ico_mod : to_Ico_mod hp a b = to_Ico_mod hp a c ↔ ∃ n : ℤ, c - b = n • p := begin refine ⟨λ h, ⟨to_Ico_div hp a c - to_Ico_div hp a b, _⟩, λ h, _⟩, From 7581030920af3dcb241d1df0e36f6ec8289dd6be Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 6 May 2023 21:28:16 +0000 Subject: [PATCH 0963/1291] =?UTF-8?q?feat(order/category):=20Free=20functo?= =?UTF-8?q?r=20`Lat=20=E2=A5=A4=20BddLat`=20(#18949)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Construct the free functor from the category of lattices to the category of bounded lattices. Concretely, it adds a bottom and a top element. Co-authored-by: Eric Wieser --- src/order/category/BddLat.lean | 39 +++++++ src/order/hom/lattice.lean | 182 +++++++++++++++++++++++++++++++++ 2 files changed, 221 insertions(+) diff --git a/src/order/category/BddLat.lean b/src/order/category/BddLat.lean index f6166c180a2a9..9ba37c0bf2829 100644 --- a/src/order/category/BddLat.lean +++ b/src/order/category/BddLat.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ +import category_theory.adjunction.opposites import order.category.BddOrd import order.category.Lat import order.category.Semilat @@ -123,3 +124,41 @@ lemma BddLat_dual_comp_forget_to_SemilatSup : lemma BddLat_dual_comp_forget_to_SemilatInf : BddLat.dual ⋙ forget₂ BddLat SemilatInf = forget₂ BddLat SemilatSup ⋙ SemilatSup.dual := rfl + +/-- The functor that adds a bottom and a top element to a lattice. This is the free functor. -/ +def Lat_to_BddLat : Lat.{u} ⥤ BddLat := +{ obj := λ X, BddLat.of $ with_top $ with_bot X, + map := λ X Y, lattice_hom.with_top_with_bot, + map_id' := λ X, lattice_hom.with_top_with_bot_id, + map_comp' := λ X Y Z _ _, lattice_hom.with_top_with_bot_comp _ _ } + +/-- `Lat_to_BddLat` is left adjoint to the forgetful functor, meaning it is the free +functor from `Lat` to `BddLat`. -/ +def Lat_to_BddLat_forget_adjunction : + Lat_to_BddLat.{u} ⊣ forget₂ BddLat Lat := +adjunction.mk_of_hom_equiv + { hom_equiv := λ X Y, + { to_fun := λ f, + { to_fun := f ∘ some ∘ some, + map_sup' := λ a b, (congr_arg f $ by refl).trans (f.map_sup' _ _), + map_inf' := λ a b, (congr_arg f $ by refl).trans (f.map_inf' _ _) }, + inv_fun := lattice_hom.with_top_with_bot', + left_inv := λ f, bounded_lattice_hom.ext $ λ a, match a with + | none := f.map_top'.symm + | some none := f.map_bot'.symm + | some (some a) := rfl + end, + right_inv := λ f, lattice_hom.ext $ λ a, rfl }, + hom_equiv_naturality_left_symm' := λ X Y Z f g, bounded_lattice_hom.ext $ λ a, match a with + | none := rfl + | some none := rfl + | some (some a) := rfl + end, + hom_equiv_naturality_right' := λ X Y Z f g, lattice_hom.ext $ λ a, rfl } + +/-- `Lat_to_BddLat` and `order_dual` commute. -/ +@[simps] def Lat_to_BddLat_comp_dual_iso_dual_comp_Lat_to_BddLat : + (Lat_to_BddLat.{u} ⋙ BddLat.dual) ≅ (Lat.dual ⋙ Lat_to_BddLat) := +adjunction.left_adjoint_uniq + (Lat_to_BddLat_forget_adjunction.comp BddLat.dual_equiv.to_adjunction) + (Lat.dual_equiv.to_adjunction.comp Lat_to_BddLat_forget_adjunction) diff --git a/src/order/hom/lattice.lean b/src/order/hom/lattice.lean index f9fd30def0bdf..771552ff68c30 100644 --- a/src/order/hom/lattice.lean +++ b/src/order/hom/lattice.lean @@ -968,3 +968,185 @@ bounded lattices. -/ (bounded_lattice_hom.dual.symm g).comp (bounded_lattice_hom.dual.symm f) := rfl end bounded_lattice_hom + +/-! ### `with_top`, `with_bot` -/ + +namespace sup_hom +variables [semilattice_sup α] [semilattice_sup β] [semilattice_sup γ] + +/-- Adjoins a `⊤` to the domain and codomain of a `sup_hom`. -/ +@[simps] protected def with_top (f : sup_hom α β) : sup_hom (with_top α) (with_top β) := +{ to_fun := option.map f, + map_sup' := λ a b, match a, b with + | ⊤, ⊤ := rfl + | ⊤, (b : α) := rfl + | (a : α), ⊤ := rfl + | (a : α), (b : α) := congr_arg _ (f.map_sup' _ _) + end } + +@[simp] lemma with_top_id : (sup_hom.id α).with_top = sup_hom.id _ := +fun_like.coe_injective option.map_id + +@[simp] lemma with_top_comp (f : sup_hom β γ) (g : sup_hom α β) : + (f.comp g).with_top = f.with_top.comp g.with_top := +fun_like.coe_injective (option.map_comp_map _ _).symm + +/-- Adjoins a `⊥` to the domain and codomain of a `sup_hom`. -/ +@[simps] protected def with_bot (f : sup_hom α β) : sup_bot_hom (with_bot α) (with_bot β) := +{ to_fun := option.map f, + map_sup' := λ a b, match a, b with + | ⊥, ⊥ := rfl + | ⊥, (b : α) := rfl + | (a : α), ⊥ := rfl + | (a : α), (b : α) := congr_arg _ (f.map_sup' _ _) + end, + map_bot' := rfl } + +@[simp] lemma with_bot_id : (sup_hom.id α).with_bot = sup_bot_hom.id _ := +fun_like.coe_injective option.map_id + +@[simp] lemma with_bot_comp (f : sup_hom β γ) (g : sup_hom α β) : + (f.comp g).with_bot = f.with_bot.comp g.with_bot := +fun_like.coe_injective (option.map_comp_map _ _).symm + +/-- Adjoins a `⊤` to the codomain of a `sup_hom`. -/ +@[simps] def with_top' [order_top β] (f : sup_hom α β) : sup_hom (with_top α) β := +{ to_fun := λ a, a.elim ⊤ f, + map_sup' := λ a b, match a, b with + | ⊤, ⊤ := top_sup_eq.symm + | ⊤, (b : α) := top_sup_eq.symm + | (a : α), ⊤ := sup_top_eq.symm + | (a : α), (b : α) := f.map_sup' _ _ + end } + +/-- Adjoins a `⊥` to the domain of a `sup_hom`. -/ +@[simps] def with_bot' [order_bot β] (f : sup_hom α β) : sup_bot_hom (with_bot α) β := +{ to_fun := λ a, a.elim ⊥ f, + map_sup' := λ a b, match a, b with + | ⊥, ⊥ := bot_sup_eq.symm + | ⊥, (b : α) := bot_sup_eq.symm + | (a : α), ⊥ := sup_bot_eq.symm + | (a : α), (b : α) := f.map_sup' _ _ + end, + map_bot' := rfl } + +end sup_hom + +namespace inf_hom +variables [semilattice_inf α] [semilattice_inf β] [semilattice_inf γ] + +/-- Adjoins a `⊤` to the domain and codomain of an `inf_hom`. -/ +@[simps] protected def with_top (f : inf_hom α β) : inf_top_hom (with_top α) (with_top β) := +{ to_fun := option.map f, + map_inf' := λ a b, match a, b with + | ⊤, ⊤ := rfl + | ⊤, (b : α) := rfl + | (a : α), ⊤ := rfl + | (a : α), (b : α) := congr_arg _ (f.map_inf' _ _) + end, + map_top' := rfl } + +@[simp] lemma with_top_id : (inf_hom.id α).with_top = inf_top_hom.id _ := +fun_like.coe_injective option.map_id + +@[simp] lemma with_top_comp (f : inf_hom β γ) (g : inf_hom α β) : + (f.comp g).with_top = f.with_top.comp g.with_top := +fun_like.coe_injective (option.map_comp_map _ _).symm + +/-- Adjoins a `⊥ to the domain and codomain of an `inf_hom`. -/ +@[simps] protected def with_bot (f : inf_hom α β) : inf_hom (with_bot α) (with_bot β) := +{ to_fun := option.map f, + map_inf' := λ a b, match a, b with + | ⊥, ⊥ := rfl + | ⊥, (b : α) := rfl + | (a : α), ⊥ := rfl + | (a : α), (b : α) := congr_arg _ (f.map_inf' _ _) + end } + +@[simp] lemma with_bot_id : (inf_hom.id α).with_bot = inf_hom.id _ := +fun_like.coe_injective option.map_id + +@[simp] lemma with_bot_comp (f : inf_hom β γ) (g : inf_hom α β) : + (f.comp g).with_bot = f.with_bot.comp g.with_bot := +fun_like.coe_injective (option.map_comp_map _ _).symm + +/-- Adjoins a `⊤` to the codomain of an `inf_hom`. -/ +@[simps] def with_top' [order_top β] (f : inf_hom α β) : inf_top_hom (with_top α) β := +{ to_fun := λ a, a.elim ⊤ f, + map_inf' := λ a b, match a, b with + | ⊤, ⊤ := top_inf_eq.symm + | ⊤, (b : α) := top_inf_eq.symm + | (a : α), ⊤ := inf_top_eq.symm + | (a : α), (b : α) := f.map_inf' _ _ + end, + map_top' := rfl } + +/-- Adjoins a `⊥` to the codomain of an `inf_hom`. -/ +@[simps] def with_bot' [order_bot β] (f : inf_hom α β) : inf_hom (with_bot α) β := +{ to_fun := λ a, a.elim ⊥ f, + map_inf' := λ a b, match a, b with + | ⊥, ⊥ := bot_inf_eq.symm + | ⊥, (b : α) := bot_inf_eq.symm + | (a : α), ⊥ := inf_bot_eq.symm + | (a : α), (b : α) := f.map_inf' _ _ + end } + +end inf_hom + +namespace lattice_hom +variables [lattice α] [lattice β] [lattice γ] + +/-- Adjoins a `⊤` to the domain and codomain of a `lattice_hom`. -/ +@[simps] protected def with_top (f : lattice_hom α β) : lattice_hom (with_top α) (with_top β) := +{ to_sup_hom := f.to_sup_hom.with_top, ..f.to_inf_hom.with_top } + +@[simp] lemma with_top_id : (lattice_hom.id α).with_top = lattice_hom.id _ := +fun_like.coe_injective option.map_id + +@[simp] lemma with_top_comp (f : lattice_hom β γ) (g : lattice_hom α β) : + (f.comp g).with_top = f.with_top.comp g.with_top := +fun_like.coe_injective (option.map_comp_map _ _).symm + +/-- Adjoins a `⊥` to the domain and codomain of a `lattice_hom`. -/ +@[simps] protected def with_bot (f : lattice_hom α β) : lattice_hom (with_bot α) (with_bot β) := +{ to_sup_hom := f.to_sup_hom.with_bot, ..f.to_inf_hom.with_bot } + +@[simp] lemma with_bot_id : (lattice_hom.id α).with_bot = lattice_hom.id _ := +fun_like.coe_injective option.map_id + +@[simp] lemma with_bot_comp (f : lattice_hom β γ) (g : lattice_hom α β) : + (f.comp g).with_bot = f.with_bot.comp g.with_bot := +fun_like.coe_injective (option.map_comp_map _ _).symm + +/-- Adjoins a `⊤` and `⊥` to the domain and codomain of a `lattice_hom`. -/ +@[simps] def with_top_with_bot (f : lattice_hom α β) : + bounded_lattice_hom (with_top $ with_bot α) (with_top $ with_bot β) := +⟨f.with_bot.with_top, rfl, rfl⟩ + +@[simp] lemma with_top_with_bot_id : + (lattice_hom.id α).with_top_with_bot = bounded_lattice_hom.id _ := +fun_like.coe_injective $ begin + refine (congr_arg option.map _).trans option.map_id, + rw with_bot_id, + refl, +end + +@[simp] lemma with_top_with_bot_comp (f : lattice_hom β γ) (g : lattice_hom α β) : + (f.comp g).with_top_with_bot = f.with_top_with_bot.comp g.with_top_with_bot := +fun_like.coe_injective $ (congr_arg option.map $ (option.map_comp_map _ _).symm).trans + (option.map_comp_map _ _).symm + +/-- Adjoins a `⊥` to the codomain of a `lattice_hom`. -/ +@[simps] def with_top' [order_top β] (f : lattice_hom α β) : lattice_hom (with_top α) β := +{ ..f.to_sup_hom.with_top', ..f.to_inf_hom.with_top' } + +/-- Adjoins a `⊥` to the domain and codomain of a `lattice_hom`. -/ +@[simps] def with_bot' [order_bot β] (f : lattice_hom α β) : lattice_hom (with_bot α) β := +{ ..f.to_sup_hom.with_bot', ..f.to_inf_hom.with_bot' } + +/-- Adjoins a `⊤` and `⊥` to the codomain of a `lattice_hom`. -/ +@[simps] def with_top_with_bot' [bounded_order β] (f : lattice_hom α β) : + bounded_lattice_hom (with_top $ with_bot α) β := +{ to_lattice_hom := f.with_bot'.with_top', map_top' := rfl, map_bot' := rfl } + +end lattice_hom From ba2eb7043b3799f585e466e4679dfb11fb1c53ec Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sun, 7 May 2023 06:38:53 +0000 Subject: [PATCH 0964/1291] feat(algebra/order/to_interval_mod): notation for `add_comm_group.modeq` (#18955) Split from #18941 --- src/algebra/order/to_interval_mod.lean | 30 ++++++++++++++------------ 1 file changed, 16 insertions(+), 14 deletions(-) diff --git a/src/algebra/order/to_interval_mod.lean b/src/algebra/order/to_interval_mod.lean index c423f022f6fc2..a8238c03d32f3 100644 --- a/src/algebra/order/to_interval_mod.lean +++ b/src/algebra/order/to_interval_mod.lean @@ -27,8 +27,8 @@ interval. * `to_Ioc_div hp a b` (where `hp : 0 < p`): The unique integer such that this multiple of `p`, subtracted from `b`, is in `Ioc a (a + p)`. * `to_Ioc_mod hp a b` (where `hp : 0 < p`): Reduce `b` to the interval `Ioc a (a + p)`. -* `add_comm_group.modeq p a b`: `a` and `b` are congruent modulo a multiple of `p`. See also - `smodeq` which is a more general version in arbitrary submodules. +* `a ≡ b [PMOD p]`: `a` and `b` are congruent modulo a multiple of `p`. See also `smodeq` which is a + more general version in arbitrary submodules. This is notation for `add_comm_group.modeq p a b`. ## TODO @@ -462,8 +462,10 @@ with `to_Ioc_mod hp a` at `b`, or `to_Ico_div hp a` disagrees with `to_Ioc_div h def modeq (p a b : α) : Prop := ∀ z : ℤ, b - z • p ∉ set.Ioo a (a + p) include hα +notation a ` ≡ `:50 b ` [PMOD `:50 p `]`:0 := modeq p a b + lemma tfae_modeq : - tfae [modeq p a b, + tfae [a ≡ b [PMOD p], to_Ico_mod hp a b ≠ to_Ioc_mod hp a b, to_Ico_mod hp a b + p = to_Ioc_mod hp a b, to_Ico_mod hp a b = a] := @@ -492,17 +494,17 @@ end variables {a b} lemma modeq_iff_to_Ico_mod_ne_to_Ioc_mod : - modeq p a b ↔ to_Ico_mod hp a b ≠ to_Ioc_mod hp a b := (tfae_modeq hp a b).out 0 1 + a ≡ b [PMOD p] ↔ to_Ico_mod hp a b ≠ to_Ioc_mod hp a b := (tfae_modeq hp a b).out 0 1 lemma modeq_iff_to_Ico_mod_add_period_eq_to_Ioc_mod : - modeq p a b ↔ to_Ico_mod hp a b + p = to_Ioc_mod hp a b := (tfae_modeq hp a b).out 0 2 + a ≡ b [PMOD p] ↔ to_Ico_mod hp a b + p = to_Ioc_mod hp a b := (tfae_modeq hp a b).out 0 2 lemma modeq_iff_to_Ico_mod_eq_left : - modeq p a b ↔ to_Ico_mod hp a b = a := (tfae_modeq hp a b).out 0 3 + a ≡ b [PMOD p] ↔ to_Ico_mod hp a b = a := (tfae_modeq hp a b).out 0 3 lemma not_modeq_iff_to_Ico_mod_eq_to_Ioc_mod : - ¬modeq p a b ↔ to_Ico_mod hp a b = to_Ioc_mod hp a b := + ¬a ≡ b [PMOD p] ↔ to_Ico_mod hp a b = to_Ioc_mod hp a b := (modeq_iff_to_Ico_mod_ne_to_Ioc_mod _).not_left -lemma modeq_iff_to_Ioc_mod_eq_right : modeq p a b ↔ to_Ioc_mod hp a b = a + p := +lemma modeq_iff_to_Ioc_mod_eq_right : a ≡ b [PMOD p] ↔ to_Ioc_mod hp a b = a + p := begin rw [modeq_iff_to_Ico_mod_ne_to_Ioc_mod hp, ne, to_Ico_mod_eq_iff hp, not_iff_comm], obtain ⟨h₁, h₂⟩ := to_Ioc_mod_mem_Ioc hp a b, @@ -511,19 +513,19 @@ begin end lemma not_modeq_iff_to_Ico_div_eq_to_Ioc_div : - ¬modeq p a b ↔ to_Ico_div hp a b = to_Ioc_div hp a b := + ¬a ≡ b [PMOD p] ↔ to_Ico_div hp a b = to_Ioc_div hp a b := by rw [not_modeq_iff_to_Ico_mod_eq_to_Ioc_mod hp, to_Ico_mod, to_Ioc_mod, sub_right_inj, (zsmul_strict_mono_left hp).injective.eq_iff] lemma modeq_iff_to_Ico_div_eq_to_Ioc_div_add_one : - modeq p a b ↔ to_Ico_div hp a b = to_Ioc_div hp a b + 1 := + a ≡ b [PMOD p] ↔ to_Ico_div hp a b = to_Ioc_div hp a b + 1 := by rw [modeq_iff_to_Ico_mod_add_period_eq_to_Ioc_mod hp, to_Ico_mod, to_Ioc_mod, ← eq_sub_iff_add_eq, sub_sub, sub_right_inj, ← add_one_zsmul, (zsmul_strict_mono_left hp).injective.eq_iff] include hp -lemma modeq_iff_eq_add_zsmul : modeq p a b ↔ ∃ z : ℤ, b = a + z • p := +lemma modeq_iff_eq_add_zsmul : a ≡ b [PMOD p] ↔ ∃ z : ℤ, b = a + z • p := begin rw [modeq_iff_to_Ico_mod_eq_left hp], split; intro h, @@ -533,16 +535,16 @@ begin exact ⟨lt_add_of_pos_right a hp, h⟩, }, end -lemma not_modeq_iff_ne_add_zsmul : ¬modeq p a b ↔ ∀ z : ℤ, b ≠ a + z • p := +lemma not_modeq_iff_ne_add_zsmul : ¬a ≡ b [PMOD p] ↔ ∀ z : ℤ, b ≠ a + z • p := by rw [modeq_iff_eq_add_zsmul hp, not_exists] lemma modeq_iff_eq_mod_zmultiples : - modeq p a b ↔ (b : α ⧸ add_subgroup.zmultiples p) = a := + a ≡ b [PMOD p] ↔ (b : α ⧸ add_subgroup.zmultiples p) = a := by simp_rw [modeq_iff_eq_add_zsmul hp, quotient_add_group.eq_iff_sub_mem, add_subgroup.mem_zmultiples_iff, eq_sub_iff_add_eq', eq_comm] lemma not_modeq_iff_ne_mod_zmultiples : - ¬modeq p a b ↔ (b : α ⧸ add_subgroup.zmultiples p) ≠ a := + ¬a ≡ b [PMOD p] ↔ (b : α ⧸ add_subgroup.zmultiples p) ≠ a := (modeq_iff_eq_mod_zmultiples hp).not end add_comm_group From b5ad141426bb005414324f89719c77c0aa3ec612 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Sun, 7 May 2023 08:46:01 +0000 Subject: [PATCH 0965/1291] chore(*): add mathlib4 synchronization comments (#18944) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.continued_fractions.computation.basic` * `algebra.continued_fractions.continuants_recurrence` * `algebra.continued_fractions.convergents_equiv` * `algebra.continued_fractions.terminated_stable` * `algebra.continued_fractions.translations` * `linear_algebra.lagrange` * `measure_theory.function.ae_measurable_sequence` * `measure_theory.measure.ae_disjoint` * `measure_theory.measure.measure_space_def` * `measure_theory.measure.null_measurable` * `ring_theory.ideal.associated_prime` --- src/algebra/continued_fractions/computation/basic.lean | 3 +++ src/algebra/continued_fractions/continuants_recurrence.lean | 3 +++ src/algebra/continued_fractions/convergents_equiv.lean | 3 +++ src/algebra/continued_fractions/terminated_stable.lean | 3 +++ src/algebra/continued_fractions/translations.lean | 3 +++ src/linear_algebra/lagrange.lean | 3 +++ src/measure_theory/function/ae_measurable_sequence.lean | 3 +++ src/measure_theory/measure/ae_disjoint.lean | 3 +++ src/measure_theory/measure/measure_space_def.lean | 3 +++ src/measure_theory/measure/null_measurable.lean | 3 +++ src/ring_theory/ideal/associated_prime.lean | 3 +++ 11 files changed, 33 insertions(+) diff --git a/src/algebra/continued_fractions/computation/basic.lean b/src/algebra/continued_fractions/computation/basic.lean index eef0b6a89ce59..a7d533da3b7d2 100644 --- a/src/algebra/continued_fractions/computation/basic.lean +++ b/src/algebra/continued_fractions/computation/basic.lean @@ -9,6 +9,9 @@ import algebra.continued_fractions.basic /-! # Computable Continued Fractions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Summary We formalise the standard computation of (regular) continued fractions for linear ordered floor diff --git a/src/algebra/continued_fractions/continuants_recurrence.lean b/src/algebra/continued_fractions/continuants_recurrence.lean index 0b18d9baff751..e3a6f7bf026d5 100644 --- a/src/algebra/continued_fractions/continuants_recurrence.lean +++ b/src/algebra/continued_fractions/continuants_recurrence.lean @@ -7,6 +7,9 @@ import algebra.continued_fractions.translations /-! # Recurrence Lemmas for the `continuants` Function of Continued Fractions. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Summary Given a generalized continued fraction `g`, for all `n ≥ 1`, we prove that the `continuants` diff --git a/src/algebra/continued_fractions/convergents_equiv.lean b/src/algebra/continued_fractions/convergents_equiv.lean index bc5d40a4f6e8b..1018b2c393850 100644 --- a/src/algebra/continued_fractions/convergents_equiv.lean +++ b/src/algebra/continued_fractions/convergents_equiv.lean @@ -11,6 +11,9 @@ import tactic.ring /-! # Equivalence of Recursive and Direct Computations of `gcf` Convergents +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Summary We show the equivalence of two computations of convergents (recurrence relation (`convergents`) vs. diff --git a/src/algebra/continued_fractions/terminated_stable.lean b/src/algebra/continued_fractions/terminated_stable.lean index 097b544df7a02..86e0c340ce6f0 100644 --- a/src/algebra/continued_fractions/terminated_stable.lean +++ b/src/algebra/continued_fractions/terminated_stable.lean @@ -7,6 +7,9 @@ import algebra.continued_fractions.translations /-! # Stabilisation of gcf Computations Under Termination +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Summary We show that the continuants and convergents of a gcf stabilise once the gcf terminates. diff --git a/src/algebra/continued_fractions/translations.lean b/src/algebra/continued_fractions/translations.lean index dc8bfb1a2c62b..ae7b5a8324b5d 100644 --- a/src/algebra/continued_fractions/translations.lean +++ b/src/algebra/continued_fractions/translations.lean @@ -7,6 +7,9 @@ import algebra.continued_fractions.basic /-! # Basic Translation Lemmas Between Functions Defined for Continued Fractions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Summary Some simple translation lemmas between the different definitions of functions defined in diff --git a/src/linear_algebra/lagrange.lean b/src/linear_algebra/lagrange.lean index 24416e95db453..1a0948723dd8f 100644 --- a/src/linear_algebra/lagrange.lean +++ b/src/linear_algebra/lagrange.lean @@ -11,6 +11,9 @@ import ring_theory.polynomial.basic /-! # Lagrange interpolation +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions * In everything that follows, `s : finset ι` is a finite set of indexes, with `v : ι → F` an indexing of the field over some type. We call the image of v on s the interpolation nodes, diff --git a/src/measure_theory/function/ae_measurable_sequence.lean b/src/measure_theory/function/ae_measurable_sequence.lean index cef9c24c6810e..90a71460a9acd 100644 --- a/src/measure_theory/function/ae_measurable_sequence.lean +++ b/src/measure_theory/function/ae_measurable_sequence.lean @@ -9,6 +9,9 @@ import measure_theory.measurable_space /-! # Sequence of measurable functions associated to a sequence of a.e.-measurable functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define here tools to prove statements about limits (infi, supr...) of sequences of `ae_measurable` functions. Given a sequence of a.e.-measurable functions `f : ι → α → β` with hypothesis diff --git a/src/measure_theory/measure/ae_disjoint.lean b/src/measure_theory/measure/ae_disjoint.lean index f9148196a88b3..6b740e0197ffc 100644 --- a/src/measure_theory/measure/ae_disjoint.lean +++ b/src/measure_theory/measure/ae_disjoint.lean @@ -8,6 +8,9 @@ import measure_theory.measure.measure_space_def /-! # Almost everywhere disjoint sets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We say that sets `s` and `t` are `μ`-a.e. disjoint (see `measure_theory.ae_disjoint`) if their intersection has measure zero. This assumption can be used instead of `disjoint` in most theorems in measure theory. diff --git a/src/measure_theory/measure/measure_space_def.lean b/src/measure_theory/measure/measure_space_def.lean index c6379cbe713a8..cb8949247984b 100644 --- a/src/measure_theory/measure/measure_space_def.lean +++ b/src/measure_theory/measure/measure_space_def.lean @@ -9,6 +9,9 @@ import order.filter.countable_Inter /-! # Measure spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines measure spaces, the almost-everywhere filter and ae_measurable functions. See `measure_theory.measure_space` for their properties and for extended documentation. diff --git a/src/measure_theory/measure/null_measurable.lean b/src/measure_theory/measure/null_measurable.lean index 3786a06400631..1cdffc8d4a321 100644 --- a/src/measure_theory/measure/null_measurable.lean +++ b/src/measure_theory/measure/null_measurable.lean @@ -8,6 +8,9 @@ import measure_theory.measure.ae_disjoint /-! # Null measurable sets and complete measures +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions ### Null measurable sets and functions diff --git a/src/ring_theory/ideal/associated_prime.lean b/src/ring_theory/ideal/associated_prime.lean index db5a3e4f96180..5306118f8f003 100644 --- a/src/ring_theory/ideal/associated_prime.lean +++ b/src/ring_theory/ideal/associated_prime.lean @@ -12,6 +12,9 @@ import ring_theory.noetherian # Associated primes of a module +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We provide the definition and related lemmas about associated primes of modules. ## Main definition From 483dd86cfec4a1380d22b1f6acd4c3dc53f501ff Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Sun, 7 May 2023 10:16:22 +0000 Subject: [PATCH 0966/1291] chore(probability/kernel/basic): split into three files (#18957) Split `probability/kernel/basic` into 3 files: - basic: definitions of a kernel, classes of kernels and some elementary kernels - measurable_integral: measurability of the Lebesgue integral against a kernel. I am about to add a lot more material to this file, about (ae-)strong measurability, the Bochner integral, etc. This planned increase in size is the motivation for this PR. - with_density: the kernel `with_density`. It requires results about measurability of the integral, hence it cannot remain in basic. --- src/probability/kernel/basic.lean | 375 +----------------- src/probability/kernel/composition.lean | 2 +- .../kernel/measurable_integral.lean | 160 ++++++++ src/probability/kernel/with_density.lean | 239 +++++++++++ 4 files changed, 418 insertions(+), 358 deletions(-) create mode 100644 src/probability/kernel/measurable_integral.lean create mode 100644 src/probability/kernel/with_density.lean diff --git a/src/probability/kernel/basic.lean b/src/probability/kernel/basic.lean index a3163ffd2284f..e4cadd40d6d8d 100644 --- a/src/probability/kernel/basic.lean +++ b/src/probability/kernel/basic.lean @@ -17,36 +17,30 @@ sets `s` of `β`, `a ↦ κ a s` is measurable. ## Main definitions Classes of kernels: -* `kernel α β`: kernels from `α` to `β`, defined as the `add_submonoid` of the measurable - functions in `α → measure β`. -* `is_markov_kernel κ`: a kernel from `α` to `β` is said to be a Markov kernel if for all `a : α`, - `k a` is a probability measure. -* `is_finite_kernel κ`: a kernel from `α` to `β` is said to be finite if there exists `C : ℝ≥0∞` - such that `C < ∞` and for all `a : α`, `κ a univ ≤ C`. This implies in particular that all - measures in the image of `κ` are finite, but is stronger since it requires an uniform bound. This - stronger condition is necessary to ensure that the composition of two finite kernels is finite. -* `is_s_finite_kernel κ`: a kernel is called s-finite if it is a countable sum of finite kernels. +* `probability_theory.kernel α β`: kernels from `α` to `β`, defined as the `add_submonoid` of the + measurable functions in `α → measure β`. +* `probability_theory.is_markov_kernel κ`: a kernel from `α` to `β` is said to be a Markov kernel + if for all `a : α`, `k a` is a probability measure. +* `probability_theory.is_finite_kernel κ`: a kernel from `α` to `β` is said to be finite if there + exists `C : ℝ≥0∞` such that `C < ∞` and for all `a : α`, `κ a univ ≤ C`. This implies in + particular that all measures in the image of `κ` are finite, but is stronger since it requires an + uniform bound. This stronger condition is necessary to ensure that the composition of two finite + kernels is finite. +* `probability_theory.kernel.is_s_finite_kernel κ`: a kernel is called s-finite if it is a countable + sum of finite kernels. Particular kernels: -* `deterministic (f : α → β) (hf : measurable f)`: kernel `a ↦ measure.dirac (f a)`. -* `const α (μβ : measure β)`: constant kernel `a ↦ μβ`. -* `kernel.restrict κ (hs : measurable_set s)`: kernel for which the image of `a : α` is - `(κ a).restrict s`. +* `probability_theory.kernel.deterministic (f : α → β) (hf : measurable f)`: + kernel `a ↦ measure.dirac (f a)`. +* `probability_theory.kernel.const α (μβ : measure β)`: constant kernel `a ↦ μβ`. +* `probability_theory.kernel.restrict κ (hs : measurable_set s)`: kernel for which the image of + `a : α` is `(κ a).restrict s`. Integral: `∫⁻ b, f b ∂(kernel.restrict κ hs a) = ∫⁻ b in s, f b ∂(κ a)` -* `kernel.with_density κ (f : α → β → ℝ≥0∞)`: kernel `a ↦ (κ a).with_density (f a)`. - It is defined if `κ` is s-finite. If `f` is finite everywhere, then this is also an s-finite - kernel. The class of s-finite kernels is the smallest class of kernels that contains finite - kernels and which is stable by `with_density`. - Integral: `∫⁻ b, g b ∂(with_density κ f a) = ∫⁻ b, f a b * g b ∂(κ a)` ## Main statements -* `ext_fun`: if `∫⁻ b, f b ∂(κ a) = ∫⁻ b, f b ∂(η a)` for all measurable functions `f` and all `a`, - then the two kernels `κ` and `η` are equal. - -* `measurable_lintegral`: the function `a ↦ ∫⁻ b, f a b ∂(κ a)` is measurable, for an s-finite - kernel `κ : kernel α β` and a function `f : α → β → ℝ≥0∞` such that `function.uncurry f` - is measurable. +* `probability_theory.kernel.ext_fun`: if `∫⁻ b, f b ∂(κ a) = ∫⁻ b, f b ∂(η a)` for all measurable + functions `f` and all `a`, then the two kernels `κ` and `η` are equal. -/ @@ -427,339 +421,6 @@ end end restrict -section measurable_lintegral - -/-- This is an auxiliary lemma for `measurable_prod_mk_mem`. -/ -lemma measurable_prod_mk_mem_of_finite (κ : kernel α β) {t : set (α × β)} (ht : measurable_set t) - (hκs : ∀ a, is_finite_measure (κ a)) : - measurable (λ a, κ a {b | (a, b) ∈ t}) := -begin - -- `t` is a measurable set in the product `α × β`: we use that the product σ-algebra is generated - -- by boxes to prove the result by induction. - refine measurable_space.induction_on_inter generate_from_prod.symm is_pi_system_prod _ _ _ _ ht, - { -- case `t = ∅` - simp only [set.mem_empty_iff_false, set.set_of_false, measure_empty, measurable_const], }, - { -- case of a box: `t = t₁ ×ˢ t₂` for measurable sets `t₁` and `t₂` - intros t' ht', - simp only [set.mem_image2, set.mem_set_of_eq, exists_and_distrib_left] at ht', - obtain ⟨t₁, ht₁, t₂, ht₂, rfl⟩ := ht', - simp only [set.prod_mk_mem_set_prod_eq], - classical, - have h_eq_ite : (λ a, κ a {b : β | a ∈ t₁ ∧ b ∈ t₂}) = λ a, ite (a ∈ t₁) (κ a t₂) 0, - { ext1 a, - split_ifs, - { simp only [h, true_and], refl, }, - { simp only [h, false_and, set.set_of_false, set.inter_empty, measure_empty], }, }, - rw h_eq_ite, - exact measurable.ite ht₁ (kernel.measurable_coe κ ht₂) measurable_const }, - { -- we assume that the result is true for `t` and we prove it for `tᶜ` - intros t' ht' h_meas, - have h_eq_sdiff : ∀ a, {b : β | (a, b) ∈ t'ᶜ} = set.univ \ {b : β | (a, b) ∈ t'}, - { intro a, - ext1 b, - simp only [set.mem_compl_iff, set.mem_set_of_eq, set.mem_diff, set.mem_univ, true_and], }, - simp_rw h_eq_sdiff, - have : (λ a, κ a (set.univ \ {b : β | (a, b) ∈ t'})) - = (λ a, (κ a set.univ - κ a {b : β | (a, b) ∈ t'})), - { ext1 a, - rw [← set.diff_inter_self_eq_diff, set.inter_univ, measure_diff], - { exact set.subset_univ _, }, - { exact (@measurable_prod_mk_left α β _ _ a) t' ht', }, - { exact measure_ne_top _ _, }, }, - rw this, - exact measurable.sub (kernel.measurable_coe κ measurable_set.univ) h_meas, }, - { -- we assume that the result is true for a family of disjoint sets and prove it for their union - intros f h_disj hf_meas hf, - have h_Union : (λ a, κ a {b : β | (a, b) ∈ ⋃ i, f i}) = λ a, κ a (⋃ i, {b : β | (a, b) ∈ f i}), - { ext1 a, - congr' with b, - simp only [set.mem_Union, set.supr_eq_Union, set.mem_set_of_eq], - refl, }, - rw h_Union, - have h_tsum : (λ a, κ a (⋃ i, {b : β | (a, b) ∈ f i})) = λ a, ∑' i, κ a {b : β | (a, b) ∈ f i}, - { ext1 a, - rw measure_Union, - { intros i j hij s hsi hsj b hbs, - have habi : {(a, b)} ⊆ f i, by { rw set.singleton_subset_iff, exact hsi hbs, }, - have habj : {(a, b)} ⊆ f j, by { rw set.singleton_subset_iff, exact hsj hbs, }, - simpa only [set.bot_eq_empty, set.le_eq_subset, set.singleton_subset_iff, - set.mem_empty_iff_false] using h_disj hij habi habj, }, - { exact λ i, (@measurable_prod_mk_left α β _ _ a) _ (hf_meas i), }, }, - rw h_tsum, - exact measurable.ennreal_tsum hf, }, -end - -lemma measurable_prod_mk_mem (κ : kernel α β) [is_s_finite_kernel κ] - {t : set (α × β)} (ht : measurable_set t) : - measurable (λ a, κ a {b | (a, b) ∈ t}) := -begin - rw ← kernel_sum_seq κ, - have : ∀ a, kernel.sum (seq κ) a {b : β | (a, b) ∈ t} = ∑' n, seq κ n a {b : β | (a, b) ∈ t}, - from λ a, kernel.sum_apply' _ _ (measurable_prod_mk_left ht), - simp_rw this, - refine measurable.ennreal_tsum (λ n, _), - exact measurable_prod_mk_mem_of_finite (seq κ n) ht infer_instance, -end - -lemma measurable_lintegral_indicator_const (κ : kernel α β) [is_s_finite_kernel κ] - {t : set (α × β)} (ht : measurable_set t) (c : ℝ≥0∞) : - measurable (λ a, ∫⁻ b, t.indicator (function.const (α × β) c) (a, b) ∂(κ a)) := -begin - simp_rw lintegral_indicator_const_comp measurable_prod_mk_left ht _, - exact measurable.const_mul (measurable_prod_mk_mem _ ht) c, -end - -/-- For an s-finite kernel `κ` and a function `f : α → β → ℝ≥0∞` which is measurable when seen as a -map from `α × β` (hypothesis `measurable (function.uncurry f)`), the integral -`a ↦ ∫⁻ b, f a b ∂(κ a)` is measurable. -/ -theorem measurable_lintegral (κ : kernel α β) [is_s_finite_kernel κ] - {f : α → β → ℝ≥0∞} (hf : measurable (function.uncurry f)) : - measurable (λ a, ∫⁻ b, f a b ∂(κ a)) := -begin - let F : ℕ → simple_func (α × β) ℝ≥0∞ := simple_func.eapprox (function.uncurry f), - have h : ∀ a, (⨆ n, F n a) = function.uncurry f a, - from simple_func.supr_eapprox_apply (function.uncurry f) hf, - simp only [prod.forall, function.uncurry_apply_pair] at h, - simp_rw ← h, - have : ∀ a, ∫⁻ b, (⨆ n, F n (a, b)) ∂(κ a) = ⨆ n, ∫⁻ b, F n (a, b) ∂(κ a), - { intro a, - rw lintegral_supr, - { exact λ n, (F n).measurable.comp measurable_prod_mk_left, }, - { exact λ i j hij b, simple_func.monotone_eapprox (function.uncurry f) hij _, }, }, - simp_rw this, - refine measurable_supr (λ n, simple_func.induction _ _ (F n)), - { intros c t ht, - simp only [simple_func.const_zero, simple_func.coe_piecewise, simple_func.coe_const, - simple_func.coe_zero, set.piecewise_eq_indicator], - exact measurable_lintegral_indicator_const κ ht c, }, - { intros g₁ g₂ h_disj hm₁ hm₂, - simp only [simple_func.coe_add, pi.add_apply], - have h_add : (λ a, ∫⁻ b, g₁ (a, b) + g₂ (a, b) ∂(κ a)) - = (λ a, ∫⁻ b, g₁ (a, b) ∂(κ a)) + (λ a, ∫⁻ b, g₂ (a, b) ∂(κ a)), - { ext1 a, - rw [pi.add_apply, lintegral_add_left (g₁.measurable.comp measurable_prod_mk_left)], }, - rw h_add, - exact measurable.add hm₁ hm₂, }, -end - -lemma measurable_lintegral' (κ : kernel α β) [is_s_finite_kernel κ] - {f : β → ℝ≥0∞} (hf : measurable f) : - measurable (λ a, ∫⁻ b, f b ∂(κ a)) := -measurable_lintegral κ (hf.comp measurable_snd) - -lemma measurable_set_lintegral (κ : kernel α β) [is_s_finite_kernel κ] - {f : α → β → ℝ≥0∞} (hf : measurable (function.uncurry f)) {s : set β} (hs : measurable_set s) : - measurable (λ a, ∫⁻ b in s, f a b ∂(κ a)) := -by { simp_rw ← lintegral_restrict κ hs, exact measurable_lintegral _ hf } - -lemma measurable_set_lintegral' (κ : kernel α β) [is_s_finite_kernel κ] - {f : β → ℝ≥0∞} (hf : measurable f) {s : set β} (hs : measurable_set s) : - measurable (λ a, ∫⁻ b in s, f b ∂(κ a)) := -measurable_set_lintegral κ (hf.comp measurable_snd) hs - -end measurable_lintegral - -section with_density -variables {f : α → β → ℝ≥0∞} - -/-- Kernel with image `(κ a).with_density (f a)` if `function.uncurry f` is measurable, and -with image 0 otherwise. If `function.uncurry f` is measurable, it satisfies -`∫⁻ b, g b ∂(with_density κ f hf a) = ∫⁻ b, f a b * g b ∂(κ a)`. -/ -noncomputable -def with_density (κ : kernel α β) [is_s_finite_kernel κ] (f : α → β → ℝ≥0∞) : - kernel α β := -@dite _ (measurable (function.uncurry f)) (classical.dec _) - (λ hf, ({ val := λ a, (κ a).with_density (f a), - property := - begin - refine measure.measurable_of_measurable_coe _ (λ s hs, _), - simp_rw with_density_apply _ hs, - exact measurable_set_lintegral κ hf hs, - end, } : kernel α β)) - (λ hf, 0) - -lemma with_density_of_not_measurable (κ : kernel α β) [is_s_finite_kernel κ] - (hf : ¬ measurable (function.uncurry f)) : - with_density κ f = 0 := -by { classical, exact dif_neg hf, } - -protected lemma with_density_apply (κ : kernel α β) [is_s_finite_kernel κ] - (hf : measurable (function.uncurry f)) (a : α) : - with_density κ f a = (κ a).with_density (f a) := -by { classical, rw [with_density, dif_pos hf], refl, } - -lemma with_density_apply' (κ : kernel α β) [is_s_finite_kernel κ] - (hf : measurable (function.uncurry f)) (a : α) {s : set β} (hs : measurable_set s) : - with_density κ f a s = ∫⁻ b in s, f a b ∂(κ a) := -by rw [kernel.with_density_apply κ hf, with_density_apply _ hs] - -lemma lintegral_with_density (κ : kernel α β) [is_s_finite_kernel κ] - (hf : measurable (function.uncurry f)) (a : α) {g : β → ℝ≥0∞} (hg : measurable g) : - ∫⁻ b, g b ∂(with_density κ f a) = ∫⁻ b, f a b * g b ∂(κ a) := -begin - rw [kernel.with_density_apply _ hf, - lintegral_with_density_eq_lintegral_mul _ (measurable.of_uncurry_left hf) hg], - simp_rw pi.mul_apply, -end - -lemma with_density_add_left (κ η : kernel α β) [is_s_finite_kernel κ] [is_s_finite_kernel η] - (f : α → β → ℝ≥0∞) : - with_density (κ + η) f = with_density κ f + with_density η f := -begin - by_cases hf : measurable (function.uncurry f), - { ext a s hs : 2, - simp only [kernel.with_density_apply _ hf, coe_fn_add, pi.add_apply, with_density_add_measure, - measure.add_apply], }, - { simp_rw [with_density_of_not_measurable _ hf], - rw zero_add, }, -end - -lemma with_density_kernel_sum [countable ι] (κ : ι → kernel α β) - (hκ : ∀ i, is_s_finite_kernel (κ i)) (f : α → β → ℝ≥0∞) : - @with_density _ _ _ _ (kernel.sum κ) (is_s_finite_kernel_sum hκ) f - = kernel.sum (λ i, with_density (κ i) f) := -begin - by_cases hf : measurable (function.uncurry f), - { ext1 a, - simp_rw [sum_apply, kernel.with_density_apply _ hf, sum_apply, - with_density_sum (λ n, κ n a) (f a)], }, - { simp_rw [with_density_of_not_measurable _ hf], - exact sum_zero.symm, }, -end - -lemma with_density_tsum [countable ι] (κ : kernel α β) [is_s_finite_kernel κ] - {f : ι → α → β → ℝ≥0∞} (hf : ∀ i, measurable (function.uncurry (f i))) : - with_density κ (∑' n, f n) = kernel.sum (λ n, with_density κ (f n)) := -begin - have h_sum_a : ∀ a, summable (λ n, f n a) := λ a, pi.summable.mpr (λ b, ennreal.summable), - have h_sum : summable (λ n, f n) := pi.summable.mpr h_sum_a, - ext a s hs : 2, - rw [sum_apply' _ a hs, with_density_apply' κ _ a hs], - swap, - { have : function.uncurry (∑' n, f n) = ∑' n, function.uncurry (f n), - { ext1 p, - simp only [function.uncurry_def], - rw [tsum_apply h_sum, tsum_apply (h_sum_a _), tsum_apply], - exact pi.summable.mpr (λ p, ennreal.summable), }, - rw this, - exact measurable.ennreal_tsum' hf, }, - have : ∫⁻ b in s, (∑' n, f n) a b ∂(κ a) = ∫⁻ b in s, (∑' n, (λ b, f n a b) b) ∂(κ a), - { congr' with b, - rw [tsum_apply h_sum, tsum_apply (h_sum_a a)], }, - rw [this, lintegral_tsum (λ n, (measurable.of_uncurry_left (hf n)).ae_measurable)], - congr' with n, - rw with_density_apply' _ (hf n) a hs, -end - -/-- If a kernel `κ` is finite and a function `f : α → β → ℝ≥0∞` is bounded, then `with_density κ f` -is finite. -/ -lemma is_finite_kernel_with_density_of_bounded (κ : kernel α β) [is_finite_kernel κ] - {B : ℝ≥0∞} (hB_top : B ≠ ∞) (hf_B : ∀ a b, f a b ≤ B) : - is_finite_kernel (with_density κ f) := -begin - by_cases hf : measurable (function.uncurry f), - { exact - ⟨⟨B * is_finite_kernel.bound κ, ennreal.mul_lt_top hB_top (is_finite_kernel.bound_ne_top κ), - λ a, - begin - rw with_density_apply' κ hf a measurable_set.univ, - calc ∫⁻ b in set.univ, f a b ∂(κ a) - ≤ ∫⁻ b in set.univ, B ∂(κ a) : lintegral_mono (hf_B a) - ... = B * κ a set.univ : by simp only [measure.restrict_univ, lintegral_const] - ... ≤ B * is_finite_kernel.bound κ : - mul_le_mul_left' (measure_le_bound κ a set.univ) _, - end⟩⟩, }, - { rw with_density_of_not_measurable _ hf, - apply_instance, }, -end - -/-- Auxiliary lemma for `is_s_finite_kernel_with_density`. -If a kernel `κ` is finite, then `with_density κ f` is s-finite. -/ -lemma is_s_finite_kernel_with_density_of_is_finite_kernel (κ : kernel α β) [is_finite_kernel κ] - (hf_ne_top : ∀ a b, f a b ≠ ∞) : - is_s_finite_kernel (with_density κ f) := -begin - -- We already have that for `f` bounded from above and a `κ` a finite kernel, - -- `with_density κ f` is finite. We write any function as a countable sum of bounded - -- functions, and decompose an s-finite kernel as a sum of finite kernels. We then use that - -- `with_density` commutes with sums for both arguments and get a sum of finite kernels. - by_cases hf : measurable (function.uncurry f), - swap, { rw with_density_of_not_measurable _ hf, apply_instance, }, - let fs : ℕ → α → β → ℝ≥0∞ := λ n a b, min (f a b) (n + 1) - min (f a b) n, - have h_le : ∀ a b n, ⌈(f a b).to_real⌉₊ ≤ n → f a b ≤ n, - { intros a b n hn, - have : (f a b).to_real ≤ n := nat.le_of_ceil_le hn, - rw ← ennreal.le_of_real_iff_to_real_le (hf_ne_top a b) _ at this, - { refine this.trans (le_of_eq _), - rw ennreal.of_real_coe_nat, }, - { norm_cast, - exact zero_le _, }, }, - have h_zero : ∀ a b n, ⌈(f a b).to_real⌉₊ ≤ n → fs n a b = 0, - { intros a b n hn, - suffices : min (f a b) (n + 1) = f a b ∧ min (f a b) n = f a b, - { simp_rw [fs, this.1, this.2, tsub_self (f a b)], }, - exact ⟨min_eq_left ((h_le a b n hn).trans (le_add_of_nonneg_right zero_le_one)), - min_eq_left (h_le a b n hn)⟩, }, - have hf_eq_tsum : f = ∑' n, fs n, - { have h_sum_a : ∀ a, summable (λ n, fs n a), - { refine λ a, pi.summable.mpr (λ b, _), - suffices : ∀ n, n ∉ finset.range ⌈(f a b).to_real⌉₊ → fs n a b = 0, - from summable_of_ne_finset_zero this, - intros n hn_not_mem, - rw [finset.mem_range, not_lt] at hn_not_mem, - exact h_zero a b n hn_not_mem, }, - ext a b : 2, - rw [tsum_apply (pi.summable.mpr h_sum_a), tsum_apply (h_sum_a a), - ennreal.tsum_eq_liminf_sum_nat], - have h_finset_sum : ∀ n, ∑ i in finset.range n, fs i a b = min (f a b) n, - { intros n, - induction n with n hn, - { simp only [finset.range_zero, finset.sum_empty, algebra_map.coe_zero, min_zero], }, - rw [finset.sum_range_succ, hn], - simp_rw [fs], - norm_cast, - rw add_tsub_cancel_iff_le, - refine min_le_min le_rfl _, - norm_cast, - exact nat.le_succ n, }, - simp_rw h_finset_sum, - refine (filter.tendsto.liminf_eq _).symm, - refine filter.tendsto.congr' _ tendsto_const_nhds, - rw [filter.eventually_eq, filter.eventually_at_top], - exact ⟨⌈(f a b).to_real⌉₊, λ n hn, (min_eq_left (h_le a b n hn)).symm⟩, }, - rw [hf_eq_tsum, with_density_tsum _ (λ (n : ℕ), _)], - swap, { exact (hf.min measurable_const).sub (hf.min measurable_const), }, - refine is_s_finite_kernel_sum (λ n, _), - suffices : is_finite_kernel (with_density κ (fs n)), by { haveI := this, apply_instance, }, - refine is_finite_kernel_with_density_of_bounded _ (ennreal.coe_ne_top : (↑n + 1) ≠ ∞) (λ a b, _), - norm_cast, - calc fs n a b ≤ min (f a b) (n + 1) : tsub_le_self - ... ≤ (n + 1) : min_le_right _ _ - ... = ↑(n + 1) : by norm_cast, -end - -/-- For a s-finite kernel `κ` and a function `f : α → β → ℝ≥0∞` which is everywhere finite, -`with_density κ f` is s-finite. -/ -theorem is_s_finite_kernel.with_density (κ : kernel α β) [is_s_finite_kernel κ] - (hf_ne_top : ∀ a b, f a b ≠ ∞) : - is_s_finite_kernel (with_density κ f) := -begin - have h_eq_sum : with_density κ f = kernel.sum (λ i, with_density (seq κ i) f), - { rw ← with_density_kernel_sum _ _, - congr, - exact (kernel_sum_seq κ).symm, }, - rw h_eq_sum, - exact is_s_finite_kernel_sum - (λ n, is_s_finite_kernel_with_density_of_is_finite_kernel (seq κ n) hf_ne_top), -end - -/-- For a s-finite kernel `κ` and a function `f : α → β → ℝ≥0`, `with_density κ f` is s-finite. -/ -instance (κ : kernel α β) [is_s_finite_kernel κ] (f : α → β → ℝ≥0) : - is_s_finite_kernel (with_density κ (λ a b, f a b)) := -is_s_finite_kernel.with_density κ (λ _ _, ennreal.coe_ne_top) - -end with_density - end kernel end probability_theory diff --git a/src/probability/kernel/composition.lean b/src/probability/kernel/composition.lean index 011bc48df906c..7177e33219350 100644 --- a/src/probability/kernel/composition.lean +++ b/src/probability/kernel/composition.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ -import probability.kernel.basic +import probability.kernel.measurable_integral /-! # Product and composition of kernels diff --git a/src/probability/kernel/measurable_integral.lean b/src/probability/kernel/measurable_integral.lean new file mode 100644 index 0000000000000..a94cedbd09d20 --- /dev/null +++ b/src/probability/kernel/measurable_integral.lean @@ -0,0 +1,160 @@ +/- +Copyright (c) 2023 Rémy Degenne. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Rémy Degenne +-/ +import probability.kernel.basic + +/-! +# Measurability of the integral against a kernel + +The Lebesgue integral of a measurable function against a kernel is measurable. + +## Main statements + +* `probability_theory.kernel.measurable_lintegral`: the function `a ↦ ∫⁻ b, f a b ∂(κ a)` is + measurable, for an s-finite kernel `κ : kernel α β` and a function `f : α → β → ℝ≥0∞` such that + `function.uncurry f` is measurable. + + +-/ + +open measure_theory probability_theory + +open_locale measure_theory ennreal nnreal big_operators + +namespace probability_theory.kernel + +variables {α β ι : Type*} {mα : measurable_space α} {mβ : measurable_space β} + +include mα mβ + +/-- This is an auxiliary lemma for `measurable_prod_mk_mem`. -/ +lemma measurable_prod_mk_mem_of_finite (κ : kernel α β) {t : set (α × β)} (ht : measurable_set t) + (hκs : ∀ a, is_finite_measure (κ a)) : + measurable (λ a, κ a {b | (a, b) ∈ t}) := +begin + -- `t` is a measurable set in the product `α × β`: we use that the product σ-algebra is generated + -- by boxes to prove the result by induction. + refine measurable_space.induction_on_inter generate_from_prod.symm is_pi_system_prod _ _ _ _ ht, + { -- case `t = ∅` + simp only [set.mem_empty_iff_false, set.set_of_false, measure_empty, measurable_const], }, + { -- case of a box: `t = t₁ ×ˢ t₂` for measurable sets `t₁` and `t₂` + intros t' ht', + simp only [set.mem_image2, set.mem_set_of_eq, exists_and_distrib_left] at ht', + obtain ⟨t₁, ht₁, t₂, ht₂, rfl⟩ := ht', + simp only [set.prod_mk_mem_set_prod_eq], + classical, + have h_eq_ite : (λ a, κ a {b : β | a ∈ t₁ ∧ b ∈ t₂}) = λ a, ite (a ∈ t₁) (κ a t₂) 0, + { ext1 a, + split_ifs, + { simp only [h, true_and], refl, }, + { simp only [h, false_and, set.set_of_false, set.inter_empty, measure_empty], }, }, + rw h_eq_ite, + exact measurable.ite ht₁ (kernel.measurable_coe κ ht₂) measurable_const }, + { -- we assume that the result is true for `t` and we prove it for `tᶜ` + intros t' ht' h_meas, + have h_eq_sdiff : ∀ a, {b : β | (a, b) ∈ t'ᶜ} = set.univ \ {b : β | (a, b) ∈ t'}, + { intro a, + ext1 b, + simp only [set.mem_compl_iff, set.mem_set_of_eq, set.mem_diff, set.mem_univ, true_and], }, + simp_rw h_eq_sdiff, + have : (λ a, κ a (set.univ \ {b : β | (a, b) ∈ t'})) + = (λ a, (κ a set.univ - κ a {b : β | (a, b) ∈ t'})), + { ext1 a, + rw [← set.diff_inter_self_eq_diff, set.inter_univ, measure_diff], + { exact set.subset_univ _, }, + { exact (@measurable_prod_mk_left α β _ _ a) t' ht', }, + { exact measure_ne_top _ _, }, }, + rw this, + exact measurable.sub (kernel.measurable_coe κ measurable_set.univ) h_meas, }, + { -- we assume that the result is true for a family of disjoint sets and prove it for their union + intros f h_disj hf_meas hf, + have h_Union : (λ a, κ a {b : β | (a, b) ∈ ⋃ i, f i}) = λ a, κ a (⋃ i, {b : β | (a, b) ∈ f i}), + { ext1 a, + congr' with b, + simp only [set.mem_Union, set.supr_eq_Union, set.mem_set_of_eq], + refl, }, + rw h_Union, + have h_tsum : (λ a, κ a (⋃ i, {b : β | (a, b) ∈ f i})) = λ a, ∑' i, κ a {b : β | (a, b) ∈ f i}, + { ext1 a, + rw measure_Union, + { intros i j hij s hsi hsj b hbs, + have habi : {(a, b)} ⊆ f i, by { rw set.singleton_subset_iff, exact hsi hbs, }, + have habj : {(a, b)} ⊆ f j, by { rw set.singleton_subset_iff, exact hsj hbs, }, + simpa only [set.bot_eq_empty, set.le_eq_subset, set.singleton_subset_iff, + set.mem_empty_iff_false] using h_disj hij habi habj, }, + { exact λ i, (@measurable_prod_mk_left α β _ _ a) _ (hf_meas i), }, }, + rw h_tsum, + exact measurable.ennreal_tsum hf, }, +end + +lemma measurable_prod_mk_mem (κ : kernel α β) [is_s_finite_kernel κ] + {t : set (α × β)} (ht : measurable_set t) : + measurable (λ a, κ a {b | (a, b) ∈ t}) := +begin + rw ← kernel_sum_seq κ, + have : ∀ a, kernel.sum (seq κ) a {b : β | (a, b) ∈ t} = ∑' n, seq κ n a {b : β | (a, b) ∈ t}, + from λ a, kernel.sum_apply' _ _ (measurable_prod_mk_left ht), + simp_rw this, + refine measurable.ennreal_tsum (λ n, _), + exact measurable_prod_mk_mem_of_finite (seq κ n) ht infer_instance, +end + +lemma measurable_lintegral_indicator_const (κ : kernel α β) [is_s_finite_kernel κ] + {t : set (α × β)} (ht : measurable_set t) (c : ℝ≥0∞) : + measurable (λ a, ∫⁻ b, t.indicator (function.const (α × β) c) (a, b) ∂(κ a)) := +begin + simp_rw lintegral_indicator_const_comp measurable_prod_mk_left ht _, + exact measurable.const_mul (measurable_prod_mk_mem _ ht) c, +end + +/-- For an s-finite kernel `κ` and a function `f : α → β → ℝ≥0∞` which is measurable when seen as a +map from `α × β` (hypothesis `measurable (function.uncurry f)`), the integral +`a ↦ ∫⁻ b, f a b ∂(κ a)` is measurable. -/ +theorem measurable_lintegral (κ : kernel α β) [is_s_finite_kernel κ] + {f : α → β → ℝ≥0∞} (hf : measurable (function.uncurry f)) : + measurable (λ a, ∫⁻ b, f a b ∂(κ a)) := +begin + let F : ℕ → simple_func (α × β) ℝ≥0∞ := simple_func.eapprox (function.uncurry f), + have h : ∀ a, (⨆ n, F n a) = function.uncurry f a, + from simple_func.supr_eapprox_apply (function.uncurry f) hf, + simp only [prod.forall, function.uncurry_apply_pair] at h, + simp_rw ← h, + have : ∀ a, ∫⁻ b, (⨆ n, F n (a, b)) ∂(κ a) = ⨆ n, ∫⁻ b, F n (a, b) ∂(κ a), + { intro a, + rw lintegral_supr, + { exact λ n, (F n).measurable.comp measurable_prod_mk_left, }, + { exact λ i j hij b, simple_func.monotone_eapprox (function.uncurry f) hij _, }, }, + simp_rw this, + refine measurable_supr (λ n, simple_func.induction _ _ (F n)), + { intros c t ht, + simp only [simple_func.const_zero, simple_func.coe_piecewise, simple_func.coe_const, + simple_func.coe_zero, set.piecewise_eq_indicator], + exact measurable_lintegral_indicator_const κ ht c, }, + { intros g₁ g₂ h_disj hm₁ hm₂, + simp only [simple_func.coe_add, pi.add_apply], + have h_add : (λ a, ∫⁻ b, g₁ (a, b) + g₂ (a, b) ∂(κ a)) + = (λ a, ∫⁻ b, g₁ (a, b) ∂(κ a)) + (λ a, ∫⁻ b, g₂ (a, b) ∂(κ a)), + { ext1 a, + rw [pi.add_apply, lintegral_add_left (g₁.measurable.comp measurable_prod_mk_left)], }, + rw h_add, + exact measurable.add hm₁ hm₂, }, +end + +lemma measurable_lintegral' (κ : kernel α β) [is_s_finite_kernel κ] + {f : β → ℝ≥0∞} (hf : measurable f) : + measurable (λ a, ∫⁻ b, f b ∂(κ a)) := +measurable_lintegral κ (hf.comp measurable_snd) + +lemma measurable_set_lintegral (κ : kernel α β) [is_s_finite_kernel κ] + {f : α → β → ℝ≥0∞} (hf : measurable (function.uncurry f)) {s : set β} (hs : measurable_set s) : + measurable (λ a, ∫⁻ b in s, f a b ∂(κ a)) := +by { simp_rw ← lintegral_restrict κ hs, exact measurable_lintegral _ hf } + +lemma measurable_set_lintegral' (κ : kernel α β) [is_s_finite_kernel κ] + {f : β → ℝ≥0∞} (hf : measurable f) {s : set β} (hs : measurable_set s) : + measurable (λ a, ∫⁻ b in s, f b ∂(κ a)) := +measurable_set_lintegral κ (hf.comp measurable_snd) hs + +end probability_theory.kernel diff --git a/src/probability/kernel/with_density.lean b/src/probability/kernel/with_density.lean new file mode 100644 index 0000000000000..e351b37fdaa48 --- /dev/null +++ b/src/probability/kernel/with_density.lean @@ -0,0 +1,239 @@ +/- +Copyright (c) 2023 Rémy Degenne. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Rémy Degenne +-/ +import probability.kernel.measurable_integral + +/-! +# With Density + +For an s-finite kernel `κ : kernel α β` and a function `f : α → β → ℝ≥0∞` which is finite +everywhere, we define `with_density κ f` as the kernel `a ↦ (κ a).with_density (f a)`. This is +an s-finite kernel. + +## Main definitions + +* `probability_theory.kernel.with_density κ (f : α → β → ℝ≥0∞)`: + kernel `a ↦ (κ a).with_density (f a)`. It is defined if `κ` is s-finite. If `f` is finite + everywhere, then this is also an s-finite kernel. The class of s-finite kernels is the smallest + class of kernels that contains finite kernels and which is stable by `with_density`. + Integral: `∫⁻ b, g b ∂(with_density κ f a) = ∫⁻ b, f a b * g b ∂(κ a)` + +## Main statements + +* `probability_theory.kernel.lintegral_with_density`: + `∫⁻ b, g b ∂(with_density κ f a) = ∫⁻ b, f a b * g b ∂(κ a)` + +-/ + +open measure_theory probability_theory + +open_locale measure_theory ennreal nnreal big_operators + +namespace probability_theory.kernel + +variables {α β ι : Type*} {mα : measurable_space α} {mβ : measurable_space β} + +include mα mβ + +variables {κ : kernel α β} {f : α → β → ℝ≥0∞} + +/-- Kernel with image `(κ a).with_density (f a)` if `function.uncurry f` is measurable, and +with image 0 otherwise. If `function.uncurry f` is measurable, it satisfies +`∫⁻ b, g b ∂(with_density κ f hf a) = ∫⁻ b, f a b * g b ∂(κ a)`. -/ +noncomputable +def with_density (κ : kernel α β) [is_s_finite_kernel κ] (f : α → β → ℝ≥0∞) : + kernel α β := +@dite _ (measurable (function.uncurry f)) (classical.dec _) + (λ hf, ({ val := λ a, (κ a).with_density (f a), + property := + begin + refine measure.measurable_of_measurable_coe _ (λ s hs, _), + simp_rw with_density_apply _ hs, + exact measurable_set_lintegral κ hf hs, + end, } : kernel α β)) + (λ hf, 0) + +lemma with_density_of_not_measurable (κ : kernel α β) [is_s_finite_kernel κ] + (hf : ¬ measurable (function.uncurry f)) : + with_density κ f = 0 := +by { classical, exact dif_neg hf, } + +protected lemma with_density_apply (κ : kernel α β) [is_s_finite_kernel κ] + (hf : measurable (function.uncurry f)) (a : α) : + with_density κ f a = (κ a).with_density (f a) := +by { classical, rw [with_density, dif_pos hf], refl, } + +lemma with_density_apply' (κ : kernel α β) [is_s_finite_kernel κ] + (hf : measurable (function.uncurry f)) (a : α) {s : set β} (hs : measurable_set s) : + with_density κ f a s = ∫⁻ b in s, f a b ∂(κ a) := +by rw [kernel.with_density_apply κ hf, with_density_apply _ hs] + +lemma lintegral_with_density (κ : kernel α β) [is_s_finite_kernel κ] + (hf : measurable (function.uncurry f)) (a : α) {g : β → ℝ≥0∞} (hg : measurable g) : + ∫⁻ b, g b ∂(with_density κ f a) = ∫⁻ b, f a b * g b ∂(κ a) := +begin + rw [kernel.with_density_apply _ hf, + lintegral_with_density_eq_lintegral_mul _ (measurable.of_uncurry_left hf) hg], + simp_rw pi.mul_apply, +end + +lemma with_density_add_left (κ η : kernel α β) [is_s_finite_kernel κ] [is_s_finite_kernel η] + (f : α → β → ℝ≥0∞) : + with_density (κ + η) f = with_density κ f + with_density η f := +begin + by_cases hf : measurable (function.uncurry f), + { ext a s hs : 2, + simp only [kernel.with_density_apply _ hf, coe_fn_add, pi.add_apply, with_density_add_measure, + measure.add_apply], }, + { simp_rw [with_density_of_not_measurable _ hf], + rw zero_add, }, +end + +lemma with_density_kernel_sum [countable ι] (κ : ι → kernel α β) + (hκ : ∀ i, is_s_finite_kernel (κ i)) (f : α → β → ℝ≥0∞) : + @with_density _ _ _ _ (kernel.sum κ) (is_s_finite_kernel_sum hκ) f + = kernel.sum (λ i, with_density (κ i) f) := +begin + by_cases hf : measurable (function.uncurry f), + { ext1 a, + simp_rw [sum_apply, kernel.with_density_apply _ hf, sum_apply, + with_density_sum (λ n, κ n a) (f a)], }, + { simp_rw [with_density_of_not_measurable _ hf], + exact sum_zero.symm, }, +end + +lemma with_density_tsum [countable ι] (κ : kernel α β) [is_s_finite_kernel κ] + {f : ι → α → β → ℝ≥0∞} (hf : ∀ i, measurable (function.uncurry (f i))) : + with_density κ (∑' n, f n) = kernel.sum (λ n, with_density κ (f n)) := +begin + have h_sum_a : ∀ a, summable (λ n, f n a) := λ a, pi.summable.mpr (λ b, ennreal.summable), + have h_sum : summable (λ n, f n) := pi.summable.mpr h_sum_a, + ext a s hs : 2, + rw [sum_apply' _ a hs, with_density_apply' κ _ a hs], + swap, + { have : function.uncurry (∑' n, f n) = ∑' n, function.uncurry (f n), + { ext1 p, + simp only [function.uncurry_def], + rw [tsum_apply h_sum, tsum_apply (h_sum_a _), tsum_apply], + exact pi.summable.mpr (λ p, ennreal.summable), }, + rw this, + exact measurable.ennreal_tsum' hf, }, + have : ∫⁻ b in s, (∑' n, f n) a b ∂(κ a) = ∫⁻ b in s, (∑' n, (λ b, f n a b) b) ∂(κ a), + { congr' with b, + rw [tsum_apply h_sum, tsum_apply (h_sum_a a)], }, + rw [this, lintegral_tsum (λ n, (measurable.of_uncurry_left (hf n)).ae_measurable)], + congr' with n, + rw with_density_apply' _ (hf n) a hs, +end + +/-- If a kernel `κ` is finite and a function `f : α → β → ℝ≥0∞` is bounded, then `with_density κ f` +is finite. -/ +lemma is_finite_kernel_with_density_of_bounded (κ : kernel α β) [is_finite_kernel κ] + {B : ℝ≥0∞} (hB_top : B ≠ ∞) (hf_B : ∀ a b, f a b ≤ B) : + is_finite_kernel (with_density κ f) := +begin + by_cases hf : measurable (function.uncurry f), + { exact + ⟨⟨B * is_finite_kernel.bound κ, ennreal.mul_lt_top hB_top (is_finite_kernel.bound_ne_top κ), + λ a, + begin + rw with_density_apply' κ hf a measurable_set.univ, + calc ∫⁻ b in set.univ, f a b ∂(κ a) + ≤ ∫⁻ b in set.univ, B ∂(κ a) : lintegral_mono (hf_B a) + ... = B * κ a set.univ : + by simp only [measure.restrict_univ, measure_theory.lintegral_const] + ... ≤ B * is_finite_kernel.bound κ : + mul_le_mul_left' (measure_le_bound κ a set.univ) _, + end⟩⟩, }, + { rw with_density_of_not_measurable _ hf, + apply_instance, }, +end + +/-- Auxiliary lemma for `is_s_finite_kernel_with_density`. +If a kernel `κ` is finite, then `with_density κ f` is s-finite. -/ +lemma is_s_finite_kernel_with_density_of_is_finite_kernel (κ : kernel α β) [is_finite_kernel κ] + (hf_ne_top : ∀ a b, f a b ≠ ∞) : + is_s_finite_kernel (with_density κ f) := +begin + -- We already have that for `f` bounded from above and a `κ` a finite kernel, + -- `with_density κ f` is finite. We write any function as a countable sum of bounded + -- functions, and decompose an s-finite kernel as a sum of finite kernels. We then use that + -- `with_density` commutes with sums for both arguments and get a sum of finite kernels. + by_cases hf : measurable (function.uncurry f), + swap, { rw with_density_of_not_measurable _ hf, apply_instance, }, + let fs : ℕ → α → β → ℝ≥0∞ := λ n a b, min (f a b) (n + 1) - min (f a b) n, + have h_le : ∀ a b n, ⌈(f a b).to_real⌉₊ ≤ n → f a b ≤ n, + { intros a b n hn, + have : (f a b).to_real ≤ n := nat.le_of_ceil_le hn, + rw ← ennreal.le_of_real_iff_to_real_le (hf_ne_top a b) _ at this, + { refine this.trans (le_of_eq _), + rw ennreal.of_real_coe_nat, }, + { norm_cast, + exact zero_le _, }, }, + have h_zero : ∀ a b n, ⌈(f a b).to_real⌉₊ ≤ n → fs n a b = 0, + { intros a b n hn, + suffices : min (f a b) (n + 1) = f a b ∧ min (f a b) n = f a b, + { simp_rw [fs, this.1, this.2, tsub_self (f a b)], }, + exact ⟨min_eq_left ((h_le a b n hn).trans (le_add_of_nonneg_right zero_le_one)), + min_eq_left (h_le a b n hn)⟩, }, + have hf_eq_tsum : f = ∑' n, fs n, + { have h_sum_a : ∀ a, summable (λ n, fs n a), + { refine λ a, pi.summable.mpr (λ b, _), + suffices : ∀ n, n ∉ finset.range ⌈(f a b).to_real⌉₊ → fs n a b = 0, + from summable_of_ne_finset_zero this, + intros n hn_not_mem, + rw [finset.mem_range, not_lt] at hn_not_mem, + exact h_zero a b n hn_not_mem, }, + ext a b : 2, + rw [tsum_apply (pi.summable.mpr h_sum_a), tsum_apply (h_sum_a a), + ennreal.tsum_eq_liminf_sum_nat], + have h_finset_sum : ∀ n, ∑ i in finset.range n, fs i a b = min (f a b) n, + { intros n, + induction n with n hn, + { simp only [finset.range_zero, finset.sum_empty, algebra_map.coe_zero, min_zero], }, + rw [finset.sum_range_succ, hn], + simp_rw [fs], + norm_cast, + rw add_tsub_cancel_iff_le, + refine min_le_min le_rfl _, + norm_cast, + exact nat.le_succ n, }, + simp_rw h_finset_sum, + refine (filter.tendsto.liminf_eq _).symm, + refine filter.tendsto.congr' _ tendsto_const_nhds, + rw [filter.eventually_eq, filter.eventually_at_top], + exact ⟨⌈(f a b).to_real⌉₊, λ n hn, (min_eq_left (h_le a b n hn)).symm⟩, }, + rw [hf_eq_tsum, with_density_tsum _ (λ (n : ℕ), _)], + swap, { exact (hf.min measurable_const).sub (hf.min measurable_const), }, + refine is_s_finite_kernel_sum (λ n, _), + suffices : is_finite_kernel (with_density κ (fs n)), by { haveI := this, apply_instance, }, + refine is_finite_kernel_with_density_of_bounded _ (ennreal.coe_ne_top : (↑n + 1) ≠ ∞) (λ a b, _), + norm_cast, + calc fs n a b ≤ min (f a b) (n + 1) : tsub_le_self + ... ≤ (n + 1) : min_le_right _ _ + ... = ↑(n + 1) : by norm_cast, +end + +/-- For a s-finite kernel `κ` and a function `f : α → β → ℝ≥0∞` which is everywhere finite, +`with_density κ f` is s-finite. -/ +theorem is_s_finite_kernel.with_density (κ : kernel α β) [is_s_finite_kernel κ] + (hf_ne_top : ∀ a b, f a b ≠ ∞) : + is_s_finite_kernel (with_density κ f) := +begin + have h_eq_sum : with_density κ f = kernel.sum (λ i, with_density (seq κ i) f), + { rw ← with_density_kernel_sum _ _, + congr, + exact (kernel_sum_seq κ).symm, }, + rw h_eq_sum, + exact is_s_finite_kernel_sum + (λ n, is_s_finite_kernel_with_density_of_is_finite_kernel (seq κ n) hf_ne_top), +end + +/-- For a s-finite kernel `κ` and a function `f : α → β → ℝ≥0`, `with_density κ f` is s-finite. -/ +instance (κ : kernel α β) [is_s_finite_kernel κ] (f : α → β → ℝ≥0) : + is_s_finite_kernel (with_density κ (λ a b, f a b)) := +is_s_finite_kernel.with_density κ (λ _ _, ennreal.coe_ne_top) + +end probability_theory.kernel From 08e1d8d4d989df3a6df86f385e9053ec8a372cc1 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Sun, 7 May 2023 19:11:01 +0000 Subject: [PATCH 0967/1291] feat(measure_theory/measure/stieltjes): add measure_Iic, measure_Ici, measure_univ (#18884) --- src/measure_theory/measure/stieltjes.lean | 154 ++++++++++++++++++++++ 1 file changed, 154 insertions(+) diff --git a/src/measure_theory/measure/stieltjes.lean b/src/measure_theory/measure/stieltjes.lean index 2c6a68198e461..9911b118cab92 100644 --- a/src/measure_theory/measure/stieltjes.lean +++ b/src/measure_theory/measure/stieltjes.lean @@ -22,6 +22,131 @@ a Borel measure `f.measure`. * `f.measure_Icc` and `f.measure_Ico` are analogous. -/ + +section move_this +-- this section contains lemmas that should be moved to appropriate places after the port to lean 4 + +open filter set +open_locale topology + +-- todo after the port: move to order/filter/at_top_bot +lemma exists_seq_monotone_tendsto_at_top_at_top (α : Type*) [semilattice_sup α] [nonempty α] + [(at_top : filter α).is_countably_generated] : + ∃ xs : ℕ → α, monotone xs ∧ tendsto xs at_top at_top := +begin + haveI h_ne_bot : (at_top : filter α).ne_bot := at_top_ne_bot, + obtain ⟨ys, h⟩ := exists_seq_tendsto (at_top : filter α), + let xs : ℕ → α := λ n, finset.sup' (finset.range (n + 1)) finset.nonempty_range_succ ys, + have h_mono : monotone xs, + { intros i j hij, + rw finset.sup'_le_iff, + intros k hk, + refine finset.le_sup'_of_le _ _ le_rfl, + rw finset.mem_range at hk ⊢, + exact hk.trans_le (add_le_add_right hij _), }, + refine ⟨xs, h_mono, _⟩, + { refine tendsto_at_top_at_top_of_monotone h_mono _, + have : ∀ (a : α), ∃ (n : ℕ), a ≤ ys n, + { rw tendsto_at_top_at_top at h, + intro a, + obtain ⟨i, hi⟩ := h a, + exact ⟨i, hi i le_rfl⟩, }, + intro a, + obtain ⟨i, hi⟩ := this a, + refine ⟨i, hi.trans _⟩, + refine finset.le_sup'_of_le _ _ le_rfl, + rw finset.mem_range_succ_iff, }, +end + +lemma exists_seq_antitone_tendsto_at_top_at_bot (α : Type*) [semilattice_inf α] [nonempty α] + [h2 : (at_bot : filter α).is_countably_generated] : + ∃ xs : ℕ → α, antitone xs ∧ tendsto xs at_top at_bot := +@exists_seq_monotone_tendsto_at_top_at_top αᵒᵈ _ _ h2 + +-- todo after the port: move to topology/algebra/order/monotone_convergence +lemma supr_eq_supr_subseq_of_antitone {ι₁ ι₂ α : Type*} [preorder ι₂] [complete_lattice α] + {l : filter ι₁} [l.ne_bot] {f : ι₂ → α} {φ : ι₁ → ι₂} (hf : antitone f) + (hφ : tendsto φ l at_bot) : + (⨆ i, f i) = (⨆ i, f (φ i)) := +le_antisymm + (supr_mono' (λ i, exists_imp_exists (λ j (hj : φ j ≤ i), hf hj) + (hφ.eventually $ eventually_le_at_bot i).exists)) + (supr_mono' (λ i, ⟨φ i, le_rfl⟩)) + +namespace measure_theory +-- todo after the port: move these lemmas to measure_theory/measure/measure_space? +variables {α : Type*} {mα : measurable_space α} +include mα + +lemma tendsto_measure_Ico_at_top [semilattice_sup α] [no_max_order α] + [(at_top : filter α).is_countably_generated] (μ : measure α) (a : α) : + tendsto (λ x, μ (Ico a x)) at_top (𝓝 (μ (Ici a))) := +begin + haveI : nonempty α := ⟨a⟩, + have h_mono : monotone (λ x, μ (Ico a x)) := λ i j hij, measure_mono (Ico_subset_Ico_right hij), + convert tendsto_at_top_supr h_mono, + obtain ⟨xs, hxs_mono, hxs_tendsto⟩ := exists_seq_monotone_tendsto_at_top_at_top α, + have h_Ici : Ici a = ⋃ n, Ico a (xs n), + { ext1 x, + simp only [mem_Ici, mem_Union, mem_Ico, exists_and_distrib_left, iff_self_and], + intro _, + obtain ⟨y, hxy⟩ := no_max_order.exists_gt x, + obtain ⟨n, hn⟩ := tendsto_at_top_at_top.mp hxs_tendsto y, + exact ⟨n, hxy.trans_le (hn n le_rfl)⟩, }, + rw [h_Ici, measure_Union_eq_supr, supr_eq_supr_subseq_of_monotone h_mono hxs_tendsto], + exact monotone.directed_le (λ i j hij, Ico_subset_Ico_right (hxs_mono hij)), +end + +lemma tendsto_measure_Ioc_at_bot [semilattice_inf α] [no_min_order α] + [(at_bot : filter α).is_countably_generated] (μ : measure α) (a : α) : + tendsto (λ x, μ (Ioc x a)) at_bot (𝓝 (μ (Iic a))) := +begin + haveI : nonempty α := ⟨a⟩, + have h_mono : antitone (λ x, μ (Ioc x a)) := λ i j hij, measure_mono (Ioc_subset_Ioc_left hij), + convert tendsto_at_bot_supr h_mono, + obtain ⟨xs, hxs_mono, hxs_tendsto⟩ := exists_seq_antitone_tendsto_at_top_at_bot α, + have h_Iic : Iic a = ⋃ n, Ioc (xs n) a, + { ext1 x, + simp only [mem_Iic, mem_Union, mem_Ioc, exists_and_distrib_right, iff_and_self], + intro _, + obtain ⟨y, hxy⟩ := no_min_order.exists_lt x, + obtain ⟨n, hn⟩ := tendsto_at_top_at_bot.mp hxs_tendsto y, + exact ⟨n, (hn n le_rfl).trans_lt hxy⟩, }, + rw [h_Iic, measure_Union_eq_supr, supr_eq_supr_subseq_of_antitone h_mono hxs_tendsto], + exact monotone.directed_le (λ i j hij, Ioc_subset_Ioc_left (hxs_mono hij)), +end + +lemma tendsto_measure_Iic_at_top [semilattice_sup α] [(at_top : filter α).is_countably_generated] + (μ : measure α) : + tendsto (λ x, μ (Iic x)) at_top (𝓝 (μ univ)) := +begin + casesI is_empty_or_nonempty α, + { have h1 : ∀ x : α, Iic x = ∅ := λ x, subsingleton.elim _ _, + have h2 : (univ : set α) = ∅ := subsingleton.elim _ _, + simp_rw [h1, h2], + exact tendsto_const_nhds, }, + have h_mono : monotone (λ x, μ (Iic x)) := λ i j hij, measure_mono (Iic_subset_Iic.mpr hij), + convert tendsto_at_top_supr h_mono, + obtain ⟨xs, hxs_mono, hxs_tendsto⟩ := exists_seq_monotone_tendsto_at_top_at_top α, + have h_univ : (univ : set α) = ⋃ n, Iic (xs n), + { ext1 x, + simp only [mem_univ, mem_Union, mem_Iic, true_iff], + obtain ⟨n, hn⟩ := tendsto_at_top_at_top.mp hxs_tendsto x, + exact ⟨n, hn n le_rfl⟩, }, + rw [h_univ, measure_Union_eq_supr, supr_eq_supr_subseq_of_monotone h_mono hxs_tendsto], + exact monotone.directed_le (λ i j hij, Iic_subset_Iic.mpr (hxs_mono hij)), +end + +lemma tendsto_measure_Ici_at_bot [semilattice_inf α] + [h : (at_bot : filter α).is_countably_generated] (μ : measure α) : + tendsto (λ x, μ (Ici x)) at_bot (𝓝 (μ univ)) := +@tendsto_measure_Iic_at_top αᵒᵈ _ _ h μ + +end measure_theory + +end move_this + + noncomputable theory open classical set filter function open ennreal (of_real) @@ -352,6 +477,35 @@ begin measure_union A measurable_set_Ioo, f.mono.le_left_lim hab, ← ennreal.of_real_add] } end +lemma measure_Iic {l : ℝ} (hf : tendsto f at_bot (𝓝 l)) (x : ℝ) : + f.measure (Iic x) = of_real (f x - l) := +begin + refine tendsto_nhds_unique (tendsto_measure_Ioc_at_bot _ _) _, + simp_rw measure_Ioc, + exact ennreal.tendsto_of_real (tendsto.const_sub _ hf), +end + +lemma measure_Ici {l : ℝ} (hf : tendsto f at_top (𝓝 l)) (x : ℝ) : + f.measure (Ici x) = of_real (l - left_lim f x) := +begin + refine tendsto_nhds_unique (tendsto_measure_Ico_at_top _ _) _, + simp_rw measure_Ico, + refine ennreal.tendsto_of_real (tendsto.sub_const _ _), + have h_le1 : ∀ x, f (x - 1) ≤ left_lim f x := λ x, monotone.le_left_lim f.mono (sub_one_lt x), + have h_le2 : ∀ x, left_lim f x ≤ f x := λ x, monotone.left_lim_le f.mono le_rfl, + refine tendsto_of_tendsto_of_tendsto_of_le_of_le (hf.comp _) hf h_le1 h_le2, + rw tendsto_at_top_at_top, + exact λ y, ⟨y + 1, λ z hyz, by rwa le_sub_iff_add_le⟩, +end + +lemma measure_univ {l u : ℝ} (hfl : tendsto f at_bot (𝓝 l)) (hfu : tendsto f at_top (𝓝 u)) : + f.measure univ = of_real (u - l) := +begin + refine tendsto_nhds_unique (tendsto_measure_Iic_at_top _) _, + simp_rw measure_Iic f hfl, + exact ennreal.tendsto_of_real (tendsto.sub_const hfu _), +end + instance : is_locally_finite_measure f.measure := ⟨λ x, ⟨Ioo (x-1) (x+1), Ioo_mem_nhds (by linarith) (by linarith), by simp⟩⟩ From b1c017582e9f18d8494e5c18602a8cb4a6f843ac Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sun, 7 May 2023 21:39:00 +0000 Subject: [PATCH 0968/1291] refactor(linear_algebra/dual): make `module.dual` reducible (#18963) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Otherwise Lean 4 can't apply `simp` lemmas about linear maps to elements of `module.dual`. There is no need for this to be reducible anyway, as all the instances on `dual` agree with the instances on linear maps. Also delete `basis.to_dual_equiv_symm_apply`, which stated `⇑(b.to_dual_equiv.symm) f = ⇑((linear_equiv.of_injective b.to_dual _).symm) (⇑((linear_equiv.of_top (linear_map.range b.to_dual) _).symm) f)` which was hardly helpful. Co-authored-by: Eric Wieser --- src/linear_algebra/dual.lean | 12 ++++-------- 1 file changed, 4 insertions(+), 8 deletions(-) diff --git a/src/linear_algebra/dual.lean b/src/linear_algebra/dual.lean index 133ca4ffbc02a..af61603118098 100644 --- a/src/linear_algebra/dual.lean +++ b/src/linear_algebra/dual.lean @@ -94,13 +94,7 @@ variables (R : Type*) (M : Type*) variables [comm_semiring R] [add_comm_monoid M] [module R M] /-- The dual space of an R-module M is the R-module of linear maps `M → R`. -/ -@[derive [add_comm_monoid, module R]] def dual := M →ₗ[R] R - -instance {S : Type*} [comm_ring S] {N : Type*} [add_comm_group N] [module S N] : - add_comm_group (dual S N) := linear_map.add_comm_group - -instance : linear_map_class (dual R M) R M R := -linear_map.semilinear_map_class +@[reducible] def dual := M →ₗ[R] R /-- The canonical pairing of a vector space and its algebraic dual. -/ def dual_pairing (R M) [comm_semiring R] [add_comm_monoid M] [module R M] : @@ -335,11 +329,13 @@ section finite variables [_root_.finite ι] /-- A vector space is linearly equivalent to its dual space. -/ -@[simps] def to_dual_equiv : M ≃ₗ[R] dual R M := linear_equiv.of_bijective b.to_dual ⟨ker_eq_bot.mp b.to_dual_ker, range_eq_top.mp b.to_dual_range⟩ +-- `simps` times out when generating this +@[simp] lemma to_dual_equiv_apply (m : M) : b.to_dual_equiv m = b.to_dual m := rfl + /-- Maps a basis for `V` to a basis for the dual space. -/ def dual_basis : basis ι R (dual R M) := b.map b.to_dual_equiv From 3f655f5297b030a87d641ad4e825af8d9679eb0b Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Mon, 8 May 2023 02:00:12 +0000 Subject: [PATCH 0969/1291] refactor(data/is_R_or_C,analysis/inner_product_space): review API (#18919) ## Drop `is_R_or_C.abs` and lemmas about it Use `has_norm.norm` instead. The norm is definitionally equal both to `real.abs` and `complex.abs`, so it's easier to specialize generic theorems to real numbers. Also, we don't have to convert between norm and `is_R_or_C.abs` here and there. - Drop `is_R_or_C.abs`, `is_R_or_C.norm_eq_abs`, `is_R_or_C.abs_of_nonneg`, `is_R_or_C.abs_zero`, `is_R_or_C.abs_one`, `is_R_or_C.abs_nonneg`, `is_R_or_C.abs_eq_zero`, `is_R_or_C.abs_ne_zero`, `is_R_or_C.abs_mul`, `is_R_or_C.abs_add`, `is_R_or_C.is_absolute_value`, `is_R_or_C.abs_abs`, `is_R_or_C.abs_pos`, `is_R_or_C.abs_neg`, `is_R_or_C.abs_inv`, `is_R_or_C.abs_div`, `is_R_or_C.abs_abs_sub_le_abs_sub`, `is_R_or_C.norm_sq_eq_abs`, `is_R_or_C.abs_to_real`, `is_R_or_C.continuous_abs`, `is_R_or_C.abs_to_complex`, `inner_product_space.core.abs_inner_symm`, `abs_inner_le_norm`. ## Rename/merge lemmas ### `is_R_or_C` - Rename `is_R_or_C.of_real_smul` to `is_R_or_C.real_smul_of_real`. - Merge `is_R_or_C.norm_real`, `is_R_or_C.norm_of_real`, and `is_R_or_C.abs_of_real` into `is_R_or_C.norm_of_real`. - Merge `is_R_or_C.abs_of_nat` and `is_R_or_C.abs_cast_nat` into `is_R_or_C.norm_nat_cast`, use `has_norm.norm`, make it a `simp, priority 900, is_R_or_C_simps, norm_cast` lemma. - Rename `is_R_or_C.mul_self_abs` to `is_R_or_C.mul_self_norm`, use `has_norm.norm`. - Rename `is_R_or_C.abs_two` to `is_R_or_C.norm_two`, use `has_norm.norm`. - Rename `is_R_or_C.abs_conj` to `is_R_or_C.norm_conj`, use `has_norm.norm`. - Rename `is_R_or_C.abs_re_le_abs` to `is_R_or_C.abs_re_le_norm`, use `has_norm.norm`. - Rename `is_R_or_C.abs_im_le_abs` to `is_R_or_C.abs_im_le_norm`, use `has_norm.norm`. - Rename `is_R_or_C.re_le_abs` and `is_R_or_C.im_le_abs` to `is_R_or_C.re_le_norm` and `is_R_or_C.im_le_norm`, respectively; use `has_norm.norm`. - Use `has_norm.norm` in `is_R_or_C.im_eq_zero_of_le` and `is_R_or_C.re_eq_self_of_le`. - Rename `is_R_or_C.abs_re_div_abs_le_one` and `is_R_or_C.abs_im_div_abs_le_one` to `is_R_or_C.abs_re_div_norm_le_one` and `is_R_or_C.abs_im_div_norm_le_one`, respectively; use `has_norm.norm`. - Rename `is_R_or_C.re_eq_abs_of_mul_conj` to `is_R_or_C.re_eq_norm_of_mul_conj`, use `has_norm.norm`. - Rename `is_R_or_C.abs_sq_re_add_conj` and `is_R_or_C.abs_sq_re_add_conj'` to `is_R_or_C.norm_sq_re_add_conj` and `is_R_or_C.norm_sq_re_conj_add`, respectively; use `has_norm.norm`. - Use `has_norm.norm` in all lemmas/definitions about `is_cau_seq` and `cau_seq` sequences of `is_R_or_C` numbers. - Rename `is_R_or_C.is_cau_seq_abs` to `is_R_or_C.is_cau_seq_norm`, use `has_norm.norm`. ### Inner products - Rename `inner_product_space.core.inner_mul_conj_re_abs` to `inner_product_space.core.inner_mul_symm_re_eq_norm`, use `has_norm.norm`. - Do the same in the root NS. - Rename `inner_self_re_abs` to `inner_self_re_eq_norm`, use `has_norm.norm`. - Rename `inner_self_abs_to_K` to `inner_self_norm_to_K`, use `has_norm.norm`. - Rename `abs_inner_symm` to `norm_inner_symm`, use `has_norm.norm`. - Rename `abs_inner_div_norm_mul_norm_eq_one_of_ne_zero_of_ne_zero_mul` to `norm_inner_div_norm_mul_norm_eq_one_of_ne_zero_of_ne_zero_mul`, use `has_norm.norm`. ## Add lemmas - Add `is_R_or_C.is_real_tfae` and `is_real_tfae.conj_eq_iff_im`. - Add `is_R_or_C.norm_sq_apply`. ## Change attributes - `is_R_or_C.zero_re'` is no longer a `simp` lemma - `is_R_or_C.norm_conj` is now a `simp` lemma. ## Misc - Reorder lemmas here and there to golf. Co-authored-by: Eric Wieser --- docs/100.yaml | 2 +- .../calculus/uniform_limits_deriv.lean | 4 +- src/analysis/complex/basic.lean | 8 +- src/analysis/complex/schwarz.lean | 2 +- src/analysis/convex/gauge.lean | 6 +- src/analysis/inner_product_space/basic.lean | 91 ++--- .../inner_product_space/rayleigh.lean | 2 +- .../inner_product_space/symmetric.lean | 2 +- .../locally_convex/continuous_of_bounded.lean | 9 +- src/analysis/normed_space/exponential.lean | 2 +- src/analysis/normed_space/extend.lean | 2 +- src/analysis/normed_space/is_R_or_C.lean | 2 +- src/analysis/normed_space/spectrum.lean | 3 +- src/data/is_R_or_C/basic.lean | 368 +++++++----------- src/measure_theory/function/l2_space.lean | 20 +- .../polynomial/cyclotomic/eval.lean | 2 +- 16 files changed, 197 insertions(+), 328 deletions(-) diff --git a/docs/100.yaml b/docs/100.yaml index 607e31750e8e4..c6c0079297bcf 100644 --- a/docs/100.yaml +++ b/docs/100.yaml @@ -315,7 +315,7 @@ title : The Cauchy-Schwarz Inequality decls : - inner_mul_inner_self_le - - abs_inner_le_norm + - norm_inner_le_norm author : Zhouhang Zhou 79: title : The Intermediate Value Theorem diff --git a/src/analysis/calculus/uniform_limits_deriv.lean b/src/analysis/calculus/uniform_limits_deriv.lean index ae111d0325927..cf8881f0e5a3e 100644 --- a/src/analysis/calculus/uniform_limits_deriv.lean +++ b/src/analysis/calculus/uniform_limits_deriv.lean @@ -349,8 +349,8 @@ begin -- * The `f' n` converge to `g'` at `x` conv { congr, funext, - rw [←norm_norm, ←norm_inv,←@is_R_or_C.norm_of_real 𝕜 _ _, - is_R_or_C.of_real_inv, ←norm_smul], }, + rw [← abs_norm, ← abs_inv, ← @is_R_or_C.norm_of_real 𝕜 _ _, + is_R_or_C.of_real_inv, ← norm_smul], }, rw ←tendsto_zero_iff_norm_tendsto_zero, have : (λ a : ι × E, (‖a.2 - x‖⁻¹ : 𝕜) • (g a.2 - g x - (g' x) (a.2 - x))) = (λ a : ι × E, (‖a.2 - x‖⁻¹ : 𝕜) • (g a.2 - g x - (f a.1 a.2 - f a.1 x))) + diff --git a/src/analysis/complex/basic.lean b/src/analysis/complex/basic.lean index 4a046d4ec7e65..a475d7eca7f07 100644 --- a/src/analysis/complex/basic.lean +++ b/src/analysis/complex/basic.lean @@ -321,7 +321,7 @@ section complex_order open_locale complex_order lemma eq_coe_norm_of_nonneg {z : ℂ} (hz : 0 ≤ z) : z = ↑‖z‖ := -by rw [eq_re_of_real_le hz, is_R_or_C.norm_of_real, real.norm_of_nonneg (complex.le_def.2 hz).1] +by rw [eq_re_of_real_le hz, is_R_or_C.norm_of_real, _root_.abs_of_nonneg (complex.le_def.2 hz).1] end complex_order @@ -334,16 +334,12 @@ open_locale complex_conjugate local notation `reC` := @is_R_or_C.re ℂ _ local notation `imC` := @is_R_or_C.im ℂ _ local notation `IC` := @is_R_or_C.I ℂ _ -local notation `absC` := @is_R_or_C.abs ℂ _ local notation `norm_sqC` := @is_R_or_C.norm_sq ℂ _ @[simp] lemma re_to_complex {x : ℂ} : reC x = x.re := rfl @[simp] lemma im_to_complex {x : ℂ} : imC x = x.im := rfl @[simp] lemma I_to_complex : IC = complex.I := rfl -@[simp] lemma norm_sq_to_complex {x : ℂ} : norm_sqC x = complex.norm_sq x := -by simp [is_R_or_C.norm_sq, complex.norm_sq] -@[simp] lemma abs_to_complex {x : ℂ} : absC x = complex.abs x := -by simp [is_R_or_C.abs, complex.abs] +@[simp] lemma norm_sq_to_complex {x : ℂ} : norm_sqC x = complex.norm_sq x := rfl section tsum variables {α : Type*} (𝕜 : Type*) [is_R_or_C 𝕜] diff --git a/src/analysis/complex/schwarz.lean b/src/analysis/complex/schwarz.lean index 96f06468ded00..4e72bd775073a 100644 --- a/src/analysis/complex/schwarz.lean +++ b/src/analysis/complex/schwarz.lean @@ -98,7 +98,7 @@ begin have hg₀ : ‖g‖₊ ≠ 0, by simpa only [hg'] using one_ne_zero, calc ‖dslope f c z‖ = ‖dslope (g ∘ f) c z‖ : begin - rw [g.dslope_comp, hgf, is_R_or_C.norm_of_real, norm_norm], + rw [g.dslope_comp, hgf, is_R_or_C.norm_of_real, abs_norm], exact λ _, hd.differentiable_at (ball_mem_nhds _ hR₁) end ... ≤ R₂ / R₁ : diff --git a/src/analysis/convex/gauge.lean b/src/analysis/convex/gauge.lean index 62cc0de3f37c7..ad9c43b0664dd 100644 --- a/src/analysis/convex/gauge.lean +++ b/src/analysis/convex/gauge.lean @@ -290,13 +290,11 @@ variables [is_R_or_C 𝕜] [module 𝕜 E] [is_scalar_tower ℝ 𝕜 E] lemma gauge_norm_smul (hs : balanced 𝕜 s) (r : 𝕜) (x : E) : gauge s (‖r‖ • x) = gauge s (r • x) := begin - rw @is_R_or_C.real_smul_eq_coe_smul 𝕜, - obtain rfl | hr := eq_or_ne r 0, - { simp only [norm_zero, is_R_or_C.of_real_zero] }, unfold gauge, congr' with θ, + rw @is_R_or_C.real_smul_eq_coe_smul 𝕜, refine and_congr_right (λ hθ, (hs.smul _).mem_smul_iff _), - rw [is_R_or_C.norm_of_real, norm_norm], + rw [is_R_or_C.norm_of_real, abs_norm], end /-- If `s` is balanced, then the Minkowski functional is ℂ-homogeneous. -/ diff --git a/src/analysis/inner_product_space/basic.lean b/src/analysis/inner_product_space/basic.lean index 361495b733589..5535f38f68fcd 100644 --- a/src/analysis/inner_product_space/basic.lean +++ b/src/analysis/inner_product_space/basic.lean @@ -154,7 +154,6 @@ include c local notation `⟪`x`, `y`⟫` := @inner 𝕜 F _ x y local notation `norm_sqK` := @is_R_or_C.norm_sq 𝕜 _ local notation `reK` := @is_R_or_C.re 𝕜 _ -local notation `absK` := @is_R_or_C.abs 𝕜 _ local notation `ext_iff` := @is_R_or_C.ext_iff 𝕜 _ local postfix `†`:90 := star_ring_end _ @@ -219,11 +218,8 @@ inner_self_eq_zero.not lemma inner_self_re_to_K (x : F) : (re ⟪x, x⟫ : 𝕜) = ⟪x, x⟫ := by norm_num [ext_iff, inner_self_im] -lemma abs_inner_symm (x y : F) : abs ⟪x, y⟫ = abs ⟪y, x⟫ := -by rw [←inner_conj_symm, abs_conj] - lemma norm_inner_symm (x y : F) : ‖⟪x, y⟫‖ = ‖⟪y, x⟫‖ := -by rw [is_R_or_C.norm_eq_abs, is_R_or_C.norm_eq_abs, abs_inner_symm] +by rw [←inner_conj_symm, norm_conj] lemma inner_neg_left (x y : F) : ⟪-x, y⟫ = -⟪x, y⟫ := by { rw [← neg_one_smul 𝕜 x, inner_smul_left], simp } @@ -237,8 +233,8 @@ by { simp [sub_eq_add_neg, inner_add_left, inner_neg_left] } lemma inner_sub_right (x y z : F) : ⟪x, y - z⟫ = ⟪x, y⟫ - ⟪x, z⟫ := by { simp [sub_eq_add_neg, inner_add_right, inner_neg_right] } -lemma inner_mul_conj_re_abs (x y : F) : re (⟪x, y⟫ * ⟪y, x⟫) = abs (⟪x, y⟫ * ⟪y, x⟫) := -by { rw [←inner_conj_symm, mul_comm], exact re_eq_abs_of_mul_conj (inner y x), } +lemma inner_mul_symm_re_eq_norm (x y : F) : re (⟪x, y⟫ * ⟪y, x⟫) = ‖⟪x, y⟫ * ⟪y, x⟫‖ := +by { rw [←inner_conj_symm, mul_comm], exact re_eq_norm_of_mul_conj (inner y x), } /-- Expand `inner (x + y) (x + y)` -/ lemma inner_add_add_self (x y : F) : ⟪x + y, x + y⟫ = ⟪x, x⟫ + ⟪x, y⟫ + ⟪y, x⟫ + ⟪y, y⟫ := @@ -310,7 +306,7 @@ add_group_norm.to_normed_add_comm_group neg' := λ x, by simp only [inner_neg_left, neg_neg, inner_neg_right], add_le' := λ x y, begin have h₁ : ‖⟪x, y⟫‖ ≤ ‖x‖ * ‖y‖ := norm_inner_le_norm _ _, - have h₂ : re ⟪x, y⟫ ≤ ‖⟪x, y⟫‖ := (re_le_abs _).trans_eq (is_R_or_C.norm_eq_abs _).symm, + have h₂ : re ⟪x, y⟫ ≤ ‖⟪x, y⟫‖ := re_le_norm _, have h₃ : re ⟪x, y⟫ ≤ ‖x‖ * ‖y‖ := h₂.trans h₁, have h₄ : re ⟪y, x⟫ ≤ ‖x‖ * ‖y‖ := by rwa [←inner_conj_symm, conj_re], have : ‖x + y‖ * ‖x + y‖ ≤ (‖x‖ + ‖y‖) * (‖x‖ + ‖y‖), @@ -362,8 +358,6 @@ variables [normed_add_comm_group F] [inner_product_space ℝ F] variables [dec_E : decidable_eq E] local notation `⟪`x`, `y`⟫` := @inner 𝕜 _ _ x y local notation `IK` := @is_R_or_C.I 𝕜 _ -local notation `absR` := has_abs.abs -local notation `absK` := @is_R_or_C.abs 𝕜 _ local postfix `†`:90 := star_ring_end _ export inner_product_space (norm_sq_eq_inner) @@ -478,20 +472,23 @@ inner_product_space.to_core.nonneg_re x lemma real_inner_self_nonneg {x : F} : 0 ≤ ⟪x, x⟫_ℝ := @inner_self_nonneg ℝ F _ _ _ x @[simp] lemma inner_self_re_to_K (x : E) : (re ⟪x, x⟫ : 𝕜) = ⟪x, x⟫ := -is_R_or_C.ext_iff.2 ⟨by simp only [of_real_re], by simp only [inner_self_im, of_real_im]⟩ +((is_R_or_C.is_real_tfae (⟪x, x⟫ : 𝕜)).out 2 3).2 (inner_self_im _) lemma inner_self_eq_norm_sq_to_K (x : E) : ⟪x, x⟫ = (‖x‖ ^ 2 : 𝕜) := by rw [← inner_self_re_to_K, ← norm_sq_eq_inner, of_real_pow] -lemma inner_self_re_abs (x : E) : re ⟪x, x⟫ = abs ⟪x, x⟫ := +lemma inner_self_re_eq_norm (x : E) : re ⟪x, x⟫ = ‖⟪x, x⟫‖ := begin conv_rhs { rw [←inner_self_re_to_K] }, symmetry, - exact is_R_or_C.abs_of_nonneg inner_self_nonneg, + exact norm_of_nonneg inner_self_nonneg, end -lemma inner_self_abs_to_K (x : E) : (absK ⟪x, x⟫ : 𝕜) = ⟪x, x⟫ := -by { rw [←inner_self_re_abs], exact inner_self_re_to_K _ } +lemma inner_self_norm_to_K (x : E) : (‖⟪x, x⟫‖ : 𝕜) = ⟪x, x⟫ := +by { rw [←inner_self_re_eq_norm], exact inner_self_re_to_K _ } + +lemma real_inner_self_abs (x : F) : |⟪x, x⟫_ℝ| = ⟪x, x⟫_ℝ := +@inner_self_norm_to_K ℝ F _ _ _ x @[simp] lemma inner_self_eq_zero {x : E} : ⟪x, x⟫ = 0 ↔ x = 0 := by rw [inner_self_eq_norm_sq_to_K, sq_eq_zero_iff, of_real_eq_zero, norm_eq_zero] @@ -505,11 +502,8 @@ by rw [← norm_sq_eq_inner, (sq_nonneg _).le_iff_eq, sq_eq_zero_iff, norm_eq_ze lemma real_inner_self_nonpos {x : F} : ⟪x, x⟫_ℝ ≤ 0 ↔ x = 0 := @inner_self_nonpos ℝ F _ _ _ x -lemma real_inner_self_abs (x : F) : absR ⟪x, x⟫_ℝ = ⟪x, x⟫_ℝ := -by { have h := @inner_self_abs_to_K ℝ F _ _ _ x, simpa using h } - -lemma abs_inner_symm (x y : E) : abs ⟪x, y⟫ = abs ⟪y, x⟫ := -by rw [←inner_conj_symm, abs_conj] +lemma norm_inner_symm (x y : E) : ‖⟪x, y⟫‖ = ‖⟪y, x⟫‖ := +by rw [←inner_conj_symm, norm_conj] @[simp] lemma inner_neg_left (x y : E) : ⟪-x, y⟫ = -⟪x, y⟫ := by { rw [← neg_one_smul 𝕜 x, inner_smul_left], simp } @@ -528,8 +522,8 @@ by { simp [sub_eq_add_neg, inner_add_left] } lemma inner_sub_right (x y z : E) : ⟪x, y - z⟫ = ⟪x, y⟫ - ⟪x, z⟫ := by { simp [sub_eq_add_neg, inner_add_right] } -lemma inner_mul_conj_re_abs (x y : E) : re (⟪x, y⟫ * ⟪y, x⟫) = abs (⟪x, y⟫ * ⟪y, x⟫) := -by { rw [←inner_conj_symm, mul_comm], exact re_eq_abs_of_mul_conj (inner y x), } +lemma inner_mul_symm_re_eq_norm (x y : E) : re (⟪x, y⟫ * ⟪y, x⟫) = ‖⟪x, y⟫ * ⟪y, x⟫‖ := +by { rw [←inner_conj_symm, mul_comm], exact re_eq_norm_of_mul_conj (inner y x), } /-- Expand `⟪x + y, x + y⟫` -/ lemma inner_add_add_self (x y : E) : ⟪x + y, x + y⟫ = ⟪x, x⟫ + ⟪x, y⟫ + ⟪y, x⟫ + ⟪y, y⟫ := @@ -939,14 +933,11 @@ begin exact inner_product_space.core.norm_inner_le_norm x y end -lemma abs_inner_le_norm (x y : E) : abs ⟪x, y⟫ ≤ ‖x‖ * ‖y‖ := -(norm_eq_abs ⟪x, y⟫).symm.trans_le (norm_inner_le_norm _ _) - lemma nnnorm_inner_le_nnnorm (x y : E) : ‖⟪x, y⟫‖₊ ≤ ‖x‖₊ * ‖y‖₊ := norm_inner_le_norm x y lemma re_inner_le_norm (x y : E) : re ⟪x, y⟫ ≤ ‖x‖ * ‖y‖ := -le_trans (re_le_abs (inner x y)) (abs_inner_le_norm x y) +le_trans (re_le_norm (inner x y)) (norm_inner_le_norm x y) /-- Cauchy–Schwarz inequality with norm -/ lemma abs_real_inner_le_norm (x y : F) : |⟪x, y⟫_ℝ| ≤ ‖x‖ * ‖y‖ := @@ -1322,9 +1313,9 @@ end /-- The real inner product of two vectors, divided by the product of their norms, has absolute value at most 1. -/ -lemma abs_real_inner_div_norm_mul_norm_le_one (x y : F) : absR (⟪x, y⟫_ℝ / (‖x‖ * ‖y‖)) ≤ 1 := +lemma abs_real_inner_div_norm_mul_norm_le_one (x y : F) : |⟪x, y⟫_ℝ / (‖x‖ * ‖y‖)| ≤ 1 := begin - rw [_root_.abs_div, _root_.abs_mul, abs_norm, abs_norm], + rw [abs_div, abs_mul, abs_norm, abs_norm], exact div_le_one_of_le (abs_real_inner_le_norm x y) (by positivity) end @@ -1339,36 +1330,31 @@ by rw [inner_smul_right, ←real_inner_self_eq_norm_mul_norm] /-- The inner product of a nonzero vector with a nonzero multiple of itself, divided by the product of their norms, has absolute value 1. -/ -lemma abs_inner_div_norm_mul_norm_eq_one_of_ne_zero_of_ne_zero_mul - {x : E} {r : 𝕜} (hx : x ≠ 0) (hr : r ≠ 0) : abs ⟪x, r • x⟫ / (‖x‖ * ‖r • x‖) = 1 := +lemma norm_inner_div_norm_mul_norm_eq_one_of_ne_zero_of_ne_zero_mul + {x : E} {r : 𝕜} (hx : x ≠ 0) (hr : r ≠ 0) : ‖⟪x, r • x⟫‖ / (‖x‖ * ‖r • x‖) = 1 := begin - have hx' : ‖x‖ ≠ 0 := by simp [norm_eq_zero, hx], - have hr' : abs r ≠ 0 := by simp [is_R_or_C.abs_eq_zero, hr], - rw [inner_smul_right, is_R_or_C.abs_mul, ←inner_self_re_abs, inner_self_eq_norm_mul_norm, - norm_smul], - rw [is_R_or_C.norm_eq_abs, ←mul_assoc, ←div_div, mul_div_cancel _ hx', - ←div_div, mul_comm, mul_div_cancel _ hr', div_self hx'], + have hx' : ‖x‖ ≠ 0 := by simp [hx], + have hr' : ‖r‖ ≠ 0 := by simp [hr], + rw [inner_smul_right, norm_mul, ← inner_self_re_eq_norm, inner_self_eq_norm_mul_norm, norm_smul], + rw [← mul_assoc, ← div_div, mul_div_cancel _ hx', + ← div_div, mul_comm, mul_div_cancel _ hr', div_self hx'], end /-- The inner product of a nonzero vector with a nonzero multiple of itself, divided by the product of their norms, has absolute value 1. -/ lemma abs_real_inner_div_norm_mul_norm_eq_one_of_ne_zero_of_ne_zero_mul - {x : F} {r : ℝ} (hx : x ≠ 0) (hr : r ≠ 0) : absR ⟪x, r • x⟫_ℝ / (‖x‖ * ‖r • x‖) = 1 := -begin - rw ← abs_to_real, - exact abs_inner_div_norm_mul_norm_eq_one_of_ne_zero_of_ne_zero_mul hx hr -end + {x : F} {r : ℝ} (hx : x ≠ 0) (hr : r ≠ 0) : |⟪x, r • x⟫_ℝ| / (‖x‖ * ‖r • x‖) = 1 := +norm_inner_div_norm_mul_norm_eq_one_of_ne_zero_of_ne_zero_mul hx hr /-- The inner product of a nonzero vector with a positive multiple of itself, divided by the product of their norms, has value 1. -/ lemma real_inner_div_norm_mul_norm_eq_one_of_ne_zero_of_pos_mul {x : F} {r : ℝ} (hx : x ≠ 0) (hr : 0 < r) : ⟪x, r • x⟫_ℝ / (‖x‖ * ‖r • x‖) = 1 := begin - rw [real_inner_smul_self_right, norm_smul, real.norm_eq_abs, ←mul_assoc ‖x‖, mul_comm _ (absR r), - mul_assoc, _root_.abs_of_nonneg (le_of_lt hr), div_self], - exact mul_ne_zero (ne_of_gt hr) - (λ h, hx (norm_eq_zero.1 (eq_zero_of_mul_self_eq_zero h))) + rw [real_inner_smul_self_right, norm_smul, real.norm_eq_abs, ←mul_assoc ‖x‖, mul_comm _ (|r|), + mul_assoc, abs_of_nonneg hr.le, div_self], + exact mul_ne_zero hr.ne' (mul_self_ne_zero.2 (norm_ne_zero_iff.2 hx)) end /-- The inner product of a nonzero vector with a negative multiple of @@ -1431,9 +1417,8 @@ begin refine ⟨hx₀, (norm_inner_eq_norm_iff hx₀ hy₀).1 $ eq_of_div_eq_one _⟩, simpa using h }, { rintro ⟨hx, ⟨r, ⟨hr, rfl⟩⟩⟩, - simp only [is_R_or_C.abs_div, is_R_or_C.abs_mul, is_R_or_C.abs_of_real, is_R_or_C.norm_eq_abs, - abs_norm], - exact abs_inner_div_norm_mul_norm_eq_one_of_ne_zero_of_ne_zero_mul hx hr } + simp only [norm_div, norm_mul, norm_of_real, abs_norm], + exact norm_inner_div_norm_mul_norm_eq_one_of_ne_zero_of_ne_zero_mul hx hr } end /-- The inner product of two vectors, divided by the product of their @@ -1452,7 +1437,7 @@ begin { have : x = 0 ∨ y = (⟪x, y⟫ / ⟪x, x⟫ : 𝕜) • x := ((@norm_inner_eq_norm_tfae 𝕜 _ _ _ _ x y).out 0 1).1 (by simp [h]), rw [this.resolve_left h₀, h], - simp [norm_smul, is_R_or_C.norm_eq_abs, inner_self_abs_to_K, h₀'] }, + simp [norm_smul, inner_self_norm_to_K, h₀'] }, { conv_lhs { rw [← h, inner_smul_right, inner_self_eq_norm_sq_to_K] }, field_simp [sq, mul_left_comm] } end @@ -1569,7 +1554,7 @@ begin { simp }, { refine (mul_le_mul_right (norm_pos_iff.2 h)).mp _, calc ‖x‖ * ‖x‖ = ‖(⟪x, x⟫ : 𝕜)‖ : by rw [← sq, inner_self_eq_norm_sq_to_K, norm_pow, - norm_of_real, norm_norm] + norm_of_real, abs_norm] ... ≤ ‖innerSL 𝕜 x‖ * ‖x‖ : (innerSL 𝕜 x).le_op_norm _ } end @@ -1867,12 +1852,12 @@ begin use a, intros s₁ hs₁ s₂ hs₂, rw ← finset.sum_sdiff_sub_sum_sdiff, - refine (_root_.abs_sub _ _).trans_lt _, + refine (abs_sub _ _).trans_lt _, have : ∀ i, 0 ≤ ‖f i‖ ^ 2 := λ i : ι, sq_nonneg _, simp only [finset.abs_sum_of_nonneg' this], have : ∑ i in s₁ \ s₂, ‖f i‖ ^ 2 + ∑ i in s₂ \ s₁, ‖f i‖ ^ 2 < (sqrt ε) ^ 2, - { rw [← hV.norm_sq_diff_sum, sq_lt_sq, - _root_.abs_of_nonneg (sqrt_nonneg _), _root_.abs_of_nonneg (norm_nonneg _)], + { rw [← hV.norm_sq_diff_sum, sq_lt_sq, abs_of_nonneg (sqrt_nonneg _), + abs_of_nonneg (norm_nonneg _)], exact H s₁ hs₁ s₂ hs₂ }, have hη := sq_sqrt (le_of_lt hε), linarith }, diff --git a/src/analysis/inner_product_space/rayleigh.lean b/src/analysis/inner_product_space/rayleigh.lean index 04807e8acdf04..4fc947239e77d 100644 --- a/src/analysis/inner_product_space/rayleigh.lean +++ b/src/analysis/inner_product_space/rayleigh.lean @@ -67,7 +67,7 @@ begin let c : 𝕜 := ↑‖x‖⁻¹ * r, have : c ≠ 0 := by simp [c, hx, hr.ne'], refine ⟨c • x, _, _⟩, - { field_simp [norm_smul, is_R_or_C.norm_eq_abs, abs_of_nonneg hr.le] }, + { field_simp [norm_smul, abs_of_pos hr] }, { rw T.rayleigh_smul x this, exact hxT } }, { rintros ⟨x, hx, hxT⟩, diff --git a/src/analysis/inner_product_space/symmetric.lean b/src/analysis/inner_product_space/symmetric.lean index e69fe416d7197..4033741a0e5b3 100644 --- a/src/analysis/inner_product_space/symmetric.lean +++ b/src/analysis/inner_product_space/symmetric.lean @@ -166,7 +166,7 @@ begin inner_add_right, inner_sub_left, inner_sub_right, hT x, ← inner_conj_symm x (T y)], suffices : (re ⟪T y, x⟫ : 𝕜) = ⟪T y, x⟫, { rw conj_eq_iff_re.mpr this, - ring_nf, }, + ring, }, { rw ← re_add_im ⟪T y, x⟫, simp_rw [h, mul_zero, add_zero], norm_cast, } }, diff --git a/src/analysis/locally_convex/continuous_of_bounded.lean b/src/analysis/locally_convex/continuous_of_bounded.lean index 044738c197eaa..7245d211bd25b 100644 --- a/src/analysis/locally_convex/continuous_of_bounded.lean +++ b/src/analysis/locally_convex/continuous_of_bounded.lean @@ -108,10 +108,9 @@ begin rw ←hy, refine (bE1 (n+1)).2.smul_mem _ hx, have h' : 0 < (n : ℝ) + 1 := n.cast_add_one_pos, - rw [norm_inv, ←nat.cast_one, ←nat.cast_add, is_R_or_C.norm_eq_abs, is_R_or_C.abs_cast_nat, - nat.cast_add, nat.cast_one, inv_le h' zero_lt_one], - norm_cast, - simp, }, + rw [norm_inv, ←nat.cast_one, ←nat.cast_add, is_R_or_C.norm_nat_cast, nat.cast_add, + nat.cast_one, inv_le h' zero_lt_one], + simp }, intros n hn, -- The converse direction follows from continuity of the scalar multiplication have hcont : continuous_at (λ (x : E), (n : 𝕜) • x) 0 := @@ -149,7 +148,7 @@ begin cases exists_nat_gt r with n hn, -- We now find a contradiction between `f (u n) ∉ V` and the absorbing property have h1 : r ≤ ‖(n : 𝕜')‖ := - by { rw [is_R_or_C.norm_eq_abs, is_R_or_C.abs_cast_nat], exact hn.le }, + by { rw [is_R_or_C.norm_nat_cast], exact hn.le }, have hn' : 0 < ‖(n : 𝕜')‖ := lt_of_lt_of_le hr h1, rw [norm_pos_iff, ne.def, nat.cast_eq_zero] at hn', have h'' : f (u n) ∈ V := diff --git a/src/analysis/normed_space/exponential.lean b/src/analysis/normed_space/exponential.lean index 879be3bb81e5b..c75855885cbb5 100644 --- a/src/analysis/normed_space/exponential.lean +++ b/src/analysis/normed_space/exponential.lean @@ -367,7 +367,7 @@ begin 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_inv, norm_pow, - nnreal.norm_eq, norm_eq_abs, abs_cast_nat, mul_comm, ←mul_assoc, ←div_eq_mul_inv], + nnreal.norm_eq, norm_nat_cast, 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 diff --git a/src/analysis/normed_space/extend.lean b/src/analysis/normed_space/extend.lean index ec74be9582adf..a9080f14c606e 100644 --- a/src/analysis/normed_space/extend.lean +++ b/src/analysis/normed_space/extend.lean @@ -122,7 +122,7 @@ lemma extend_to_𝕜'_apply (fr : F →L[ℝ] ℝ) (x : F) : le_antisymm (linear_map.mk_continuous_norm_le _ (norm_nonneg _) _) $ op_norm_le_bound _ (norm_nonneg _) $ λ x, calc ‖fr x‖ = ‖re (fr.extend_to_𝕜' x : 𝕜)‖ : congr_arg norm (fr.extend_to_𝕜'_apply_re x).symm - ... ≤ ‖(fr.extend_to_𝕜' x : 𝕜)‖ : (abs_re_le_abs _).trans_eq (norm_eq_abs _).symm + ... ≤ ‖(fr.extend_to_𝕜' x : 𝕜)‖ : abs_re_le_norm _ ... ≤ ‖(fr.extend_to_𝕜' : F →L[𝕜] 𝕜)‖ * ‖x‖ : le_op_norm _ _ end continuous_linear_map diff --git a/src/analysis/normed_space/is_R_or_C.lean b/src/analysis/normed_space/is_R_or_C.lean index b2941c8b4cc54..0b19875c5996e 100644 --- a/src/analysis/normed_space/is_R_or_C.lean +++ b/src/analysis/normed_space/is_R_or_C.lean @@ -46,7 +46,7 @@ lemma norm_smul_inv_norm' {r : ℝ} (r_nonneg : 0 ≤ r) {x : E} (hx : x ≠ 0) ‖(r * ‖x‖⁻¹ : 𝕜) • x‖ = r := begin have : ‖x‖ ≠ 0 := by simp [hx], - field_simp [norm_smul, is_R_or_C.norm_eq_abs, r_nonneg] with is_R_or_C_simps + field_simp [norm_smul, r_nonneg] with is_R_or_C_simps end lemma linear_map.bound_of_sphere_bound diff --git a/src/analysis/normed_space/spectrum.lean b/src/analysis/normed_space/spectrum.lean index 96e8739795f41..cab94d955eebb 100644 --- a/src/analysis/normed_space/spectrum.lean +++ b/src/analysis/normed_space/spectrum.lean @@ -478,8 +478,7 @@ begin 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_inv, is_R_or_C.norm_eq_abs, is_R_or_C.abs_cast_nat, - ←div_eq_mul_inv], + rw [norm_smul, mul_comm, norm_inv, is_R_or_C.norm_nat_cast, ← 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) }, diff --git a/src/data/is_R_or_C/basic.lean b/src/data/is_R_or_C/basic.lean index 2f8265c89a265..d3907aa627b28 100644 --- a/src/data/is_R_or_C/basic.lean +++ b/src/data/is_R_or_C/basic.lean @@ -108,24 +108,24 @@ theorem ext_iff {z w : K} : z = w ↔ re z = re w ∧ im z = im w := theorem ext {z w : K} (hre : re z = re w) (him : im z = im w) : z = w := ext_iff.2 ⟨hre, him⟩ -@[norm_cast] lemma of_real_zero : ((0 : ℝ) : K) = 0 := -by rw [of_real_alg, zero_smul] +@[norm_cast] lemma of_real_zero : ((0 : ℝ) : K) = 0 := algebra_map.coe_zero -@[simp, is_R_or_C_simps] lemma zero_re' : re (0 : K) = (0 : ℝ) := re.map_zero +@[is_R_or_C_simps] lemma zero_re' : re (0 : K) = (0 : ℝ) := map_zero re @[norm_cast] lemma of_real_one : ((1 : ℝ) : K) = 1 := map_one (algebra_map ℝ K) @[simp, is_R_or_C_simps] lemma one_re : re (1 : K) = 1 := by rw [← of_real_one, of_real_re] @[simp, is_R_or_C_simps] lemma one_im : im (1 : K) = 0 := by rw [← of_real_one, of_real_im] theorem of_real_injective : function.injective (coe : ℝ → K) := (algebra_map ℝ K).injective - -@[norm_cast] theorem of_real_inj {z w : ℝ} : (z : K) = (w : K) ↔ z = w := -algebra_map.coe_inj +@[norm_cast] theorem of_real_inj {z w : ℝ} : (z : K) = (w : K) ↔ z = w := algebra_map.coe_inj @[simp, is_R_or_C_simps] lemma bit0_re (z : K) : re (bit0 z) = bit0 (re z) := map_bit0 _ _ + @[simp, is_R_or_C_simps] lemma bit1_re (z : K) : re (bit1 z) = bit1 (re z) := by simp only [bit1, map_add, bit0_re, one_re] + @[simp, is_R_or_C_simps] lemma bit0_im (z : K) : im (bit0 z) = bit0 (im z) := map_bit0 _ _ + @[simp, is_R_or_C_simps] lemma bit1_im (z : K) : im (bit1 z) = bit0 (im z) := by simp only [bit1, map_add, bit0_im, one_im, add_zero] @@ -175,24 +175,29 @@ ring_hom.map_prod _ _ _ ring_hom.map_finsupp_prod _ f g @[simp, norm_cast, is_R_or_C_simps] -lemma of_real_smul (r x : ℝ) : r • (x : K) = (r : K) * (x : K) := -begin - simp_rw [← smul_eq_mul, of_real_alg r], - simp only [algebra.id.smul_eq_mul, one_mul, algebra.smul_mul_assoc], -end +lemma real_smul_of_real (r x : ℝ) : r • (x : K) = (r : K) * (x : K) := real_smul_eq_coe_mul _ _ @[is_R_or_C_simps] lemma of_real_mul_re (r : ℝ) (z : K) : re (↑r * z) = r * re z := by simp only [mul_re, of_real_im, zero_mul, of_real_re, sub_zero] + @[is_R_or_C_simps] lemma of_real_mul_im (r : ℝ) (z : K) : im (↑r * z) = r * (im z) := by simp only [add_zero, of_real_im, zero_mul, of_real_re, mul_im] -@[is_R_or_C_simps] lemma smul_re : ∀ (r : ℝ) (z : K), re (r • z) = r * (re z) := -λ r z, by { rw algebra.smul_def, apply of_real_mul_re } -@[is_R_or_C_simps] lemma smul_im : ∀ (r : ℝ) (z : K), im (r • z) = r * (im z) := -λ r z, by { rw algebra.smul_def, apply of_real_mul_im } +@[is_R_or_C_simps] lemma smul_re (r : ℝ) (z : K) : re (r • z) = r * (re z) := +by rw [real_smul_eq_coe_mul, of_real_mul_re] + +@[is_R_or_C_simps] lemma smul_im (r : ℝ) (z : K) : im (r • z) = r * (im z) := +by rw [real_smul_eq_coe_mul, of_real_mul_im] + +@[simp, norm_cast, is_R_or_C_simps] lemma norm_of_real (r : ℝ) : ‖(r : K)‖ = |r| := +norm_algebra_map' K r + +/-! ### Characteristic zero -/ -@[simp, is_R_or_C_simps] lemma norm_real (r : ℝ) : ‖(r : K)‖ = ‖r‖ := -by rw [is_R_or_C.of_real_alg, norm_smul, norm_one, mul_one] +/-- ℝ and ℂ are both of characteristic zero. -/ +@[priority 100] -- see Note [lower instance priority] +instance char_zero_R_or_C : char_zero K := +(ring_hom.char_zero_iff (algebra_map ℝ K).injective).1 infer_instance /-! ### The imaginary unit, `I` -/ @@ -220,36 +225,54 @@ by { rw ext_iff, simp only [of_real_im, conj_im, eq_self_iff_true, conj_re, and_ @[simp, is_R_or_C_simps] lemma conj_neg_I : conj (-I) = (I : K) := by rw [map_neg, conj_I, neg_neg] -lemma conj_eq_re_sub_im (z : K) : conj z = re z - (im z) * I := -by { rw ext_iff, simp only [add_zero, I_re, of_real_im, I_im, zero_sub, zero_mul, conj_im, - of_real_re, eq_self_iff_true, sub_zero, conj_re, mul_im, neg_inj, - and_self, mul_re, mul_zero, map_sub], } +lemma conj_eq_re_sub_im (z : K) : conj z = re z - im z * I := +(congr_arg conj (re_add_im z).symm).trans $ by rw [map_add, map_mul, conj_I, conj_of_real, + conj_of_real, mul_neg, sub_eq_add_neg] -@[is_R_or_C_simps] lemma conj_smul (r : ℝ) (z : K) : conj (r • z) = r • conj z := +theorem sub_conj (z : K) : z - conj z = 2 * im z * I := begin - simp_rw conj_eq_re_sub_im, - simp only [smul_re, smul_im, of_real_mul], - rw smul_sub, - simp_rw of_real_alg, - simp only [one_mul, algebra.smul_mul_assoc], + nth_rewrite 0 [← re_add_im z], + rw [conj_eq_re_sub_im, add_sub_sub_cancel, ← two_mul, mul_assoc] end -lemma conj_eq_iff_real {z : K} : conj z = z ↔ ∃ r : ℝ, z = (r : K) := +@[is_R_or_C_simps] lemma conj_smul (r : ℝ) (z : K) : conj (r • z) = r • conj z := +by rw [conj_eq_re_sub_im, conj_eq_re_sub_im, smul_re, smul_im, of_real_mul, of_real_mul, + real_smul_eq_coe_mul, mul_sub, mul_assoc] + +theorem add_conj (z : K) : z + conj z = 2 * re z := +calc z + conj z = re z + im z * I + (re z - im z * I) : by rw [re_add_im, conj_eq_re_sub_im] +... = 2 * re z : by rw [add_add_sub_cancel, two_mul] + +theorem re_eq_add_conj (z : K) : ↑(re z) = (z + conj z) / 2 := +by rw [add_conj, mul_div_cancel_left ((re z):K) two_ne_zero] + +theorem im_eq_conj_sub (z : K) : ↑(im z) = I * (conj z - z) / 2 := +by rw [← neg_inj, ← of_real_neg, ← I_mul_re, re_eq_add_conj, map_mul, conj_I, ← neg_div, ← mul_neg, + neg_sub, mul_sub, neg_mul, sub_eq_add_neg] + +/-- There are several equivalent ways to say that a number `z` is in fact a real number. -/ +theorem is_real_tfae (z : K) : + tfae [conj z = z, ∃ r : ℝ, (r : K) = z, ↑(re z) = z, im z = 0] := begin - split, + tfae_have : 1 → 4, + { intro h, + rw [← @of_real_inj K, im_eq_conj_sub, h, sub_self, mul_zero, zero_div, of_real_zero] }, + tfae_have : 4 → 3, { intro h, - suffices : im z = 0, - { use (re z), - rw ← add_zero (coe _), - convert (re_add_im z).symm, simp [this] }, - contrapose! h, - rw ← re_add_im z, - simp only [conj_of_real, ring_hom.map_add, ring_hom.map_mul, conj_I_ax], - rw [add_left_cancel_iff, ext_iff], - simpa [neg_eq_iff_add_eq_zero, add_self_eq_zero] }, - { rintros ⟨r, rfl⟩, apply conj_of_real } + conv_rhs { rw [← re_add_im z, h, of_real_zero, zero_mul, add_zero] } }, + tfae_have : 3 → 2, from λ h, ⟨_, h⟩, + tfae_have : 2 → 1, from λ ⟨r, hr⟩, hr ▸ conj_of_real _, + tfae_finish end +lemma conj_eq_iff_real {z : K} : conj z = z ↔ ∃ r : ℝ, z = (r : K) := +((is_real_tfae z).out 0 1).trans $ by simp only [eq_comm] + +lemma conj_eq_iff_re {z : K} : conj z = z ↔ ((re z) : K) = z := +(is_real_tfae z).out 0 2 + +lemma conj_eq_iff_im {z : K} : conj z = z ↔ im z = 0 := (is_real_tfae z).out 0 3 + @[simp] lemma star_def : (has_star.star : K → K) = conj := rfl variables (K) @@ -259,9 +282,6 @@ abbreviation conj_to_ring_equiv : K ≃+* Kᵐᵒᵖ := star_ring_equiv variables {K} -lemma conj_eq_iff_re {z : K} : conj z = z ↔ ((re z) : K) = z := -conj_eq_iff_real.trans ⟨by rintro ⟨r, rfl⟩; simp, λ h, ⟨_, h.symm⟩⟩ - /-- The norm squared function. -/ def norm_sq : K →*₀ ℝ := { to_fun := λ z, re z * re z + im z * im z, @@ -269,8 +289,10 @@ def norm_sq : K →*₀ ℝ := map_one' := by simp only [one_im, add_zero, mul_one, one_re, mul_zero], map_mul' := λ z w, by { simp only [mul_im, mul_re], ring } } -lemma norm_sq_eq_def {z : K} : ‖z‖^2 = (re z) * (re z) + (im z) * (im z) := norm_sq_eq_def_ax z -lemma norm_sq_eq_def' (z : K) : norm_sq z = ‖z‖^2 := by { rw norm_sq_eq_def, refl } +lemma norm_sq_apply (z : K) : norm_sq z = re z * re z + im z * im z := rfl + +lemma norm_sq_eq_def {z : K} : ‖z‖^2 = re z * re z + im z * im z := norm_sq_eq_def_ax z +lemma norm_sq_eq_def' (z : K) : norm_sq z = ‖z‖^2 := norm_sq_eq_def.symm @[is_R_or_C_simps] lemma norm_sq_zero : norm_sq (0 : K) = 0 := norm_sq.map_zero @[is_R_or_C_simps] lemma norm_sq_one : norm_sq (1 : K) = 1 := norm_sq.map_one @@ -288,16 +310,14 @@ by rw [lt_iff_le_and_ne, ne, eq_comm]; simp [norm_sq_nonneg] by simp only [norm_sq_eq_def', norm_neg] @[simp, is_R_or_C_simps] lemma norm_sq_conj (z : K) : norm_sq (conj z) = norm_sq z := -by simp only [norm_sq, neg_mul, monoid_with_zero_hom.coe_mk, - mul_neg, neg_neg] with is_R_or_C_simps +by simp only [norm_sq_apply, neg_mul, mul_neg, neg_neg] with is_R_or_C_simps @[simp, is_R_or_C_simps] lemma norm_sq_mul (z w : K) : norm_sq (z * w) = norm_sq z * norm_sq w := norm_sq.map_mul z w lemma norm_sq_add (z w : K) : norm_sq (z + w) = norm_sq z + norm_sq w + 2 * (re (z * conj w)) := -by { simp only [norm_sq, map_add, monoid_with_zero_hom.coe_mk, mul_neg, - sub_neg_eq_add] with is_R_or_C_simps, ring } +by { simp only [norm_sq_apply, map_add, mul_neg, sub_neg_eq_add] with is_R_or_C_simps, ring } lemma re_sq_le_norm_sq (z : K) : re z * re z ≤ norm_sq z := le_add_of_nonneg_right (mul_self_nonneg _) @@ -306,22 +326,12 @@ lemma im_sq_le_norm_sq (z : K) : im z * im z ≤ norm_sq z := le_add_of_nonneg_left (mul_self_nonneg _) theorem mul_conj (z : K) : z * conj z = ((norm_sq z) : K) := -by simp only [map_add, add_zero, ext_iff, monoid_with_zero_hom.coe_mk, - add_left_inj, mul_eq_mul_left_iff, zero_mul, add_comm, true_or, eq_self_iff_true, - mul_neg, add_right_neg, zero_add, norm_sq, mul_comm, and_self, - neg_neg, mul_zero, sub_eq_neg_add, neg_zero] with is_R_or_C_simps +by simp only [map_add, add_zero, ext_iff, add_left_inj, mul_eq_mul_left_iff, zero_mul, add_comm, + true_or, eq_self_iff_true, mul_neg, add_right_neg, zero_add, norm_sq_apply, mul_comm, + and_self, neg_neg, mul_zero, sub_eq_neg_add, neg_zero] with is_R_or_C_simps lemma conj_mul (x : K) : conj x * x = ((norm_sq x) : K) := by rw [mul_comm, mul_conj] -theorem add_conj (z : K) : z + conj z = 2 * (re z) := -by simp only [ext_iff, two_mul, map_add, add_zero, of_real_im, conj_im, of_real_re, - eq_self_iff_true, add_right_neg, conj_re, and_self] - -theorem sub_conj (z : K) : z - conj z = (2 * im z) * I := -by simp only [ext_iff, two_mul, sub_eq_add_neg, add_mul, map_add, add_zero, add_left_inj, zero_mul, - map_add_neg, eq_self_iff_true, add_right_neg, and_self, neg_neg, mul_zero, neg_zero] - with is_R_or_C_simps - lemma norm_sq_sub (z w : K) : norm_sq (z - w) = norm_sq z + norm_sq w - 2 * re (z * conj w) := by simp only [norm_sq_add, sub_eq_add_neg, ring_equiv.map_neg, mul_neg, norm_sq_neg, map_neg] @@ -373,15 +383,15 @@ map_zpow₀ (algebra_map ℝ K) r n lemma I_mul_I_of_nonzero : (I : K) ≠ 0 → (I : K) * I = -1 := I_mul_I_ax.resolve_left -@[simp, is_R_or_C_simps] lemma div_I (z : K) : z / I = -(z * I) := +@[simp, is_R_or_C_simps] lemma inv_I : (I : K)⁻¹ = -I := begin by_cases h : (I : K) = 0, { simp [h] }, - { field_simp [mul_assoc, I_mul_I_of_nonzero h] } + { field_simp [I_mul_I_of_nonzero h] } end -@[simp, is_R_or_C_simps] lemma inv_I : (I : K)⁻¹ = -I := -by field_simp +@[simp, is_R_or_C_simps] lemma div_I (z : K) : z / I = -(z * I) := +by rw [div_eq_mul_inv, inv_I, mul_neg] @[simp, is_R_or_C_simps] lemma norm_sq_inv (z : K) : norm_sq z⁻¹ = (norm_sq z)⁻¹ := map_inv₀ (@norm_sq K _) z @@ -389,8 +399,8 @@ map_inv₀ (@norm_sq K _) z @[simp, is_R_or_C_simps] lemma norm_sq_div (z w : K) : norm_sq (z / w) = norm_sq z / norm_sq w := map_div₀ (@norm_sq K _) z w -@[is_R_or_C_simps] lemma norm_conj {z : K} : ‖conj z‖ = ‖z‖ := -by simp only [←sqrt_norm_sq_eq_norm, norm_sq_conj] +@[simp, is_R_or_C_simps] lemma norm_conj {z : K} : ‖conj z‖ = ‖z‖ := +by simp only [← sqrt_norm_sq_eq_norm, norm_sq_conj] @[priority 100] instance : cstar_ring K := { norm_star_mul_self := λ x, (norm_mul _ _).trans $ congr_arg (* ‖x‖) norm_conj } @@ -426,194 +436,92 @@ by rw [← of_real_rat_cast, of_real_re] @[simp, is_R_or_C_simps, norm_cast] lemma rat_cast_im (q : ℚ) : im (q : K) = 0 := by rw [← of_real_rat_cast, of_real_im] -/-! ### Characteristic zero -/ -/-- ℝ and ℂ are both of characteristic zero. -/ -@[priority 100] -- see Note [lower instance priority] -instance char_zero_R_or_C : char_zero K := -char_zero_of_inj_zero $ λ n h, -by rwa [← of_real_nat_cast, of_real_eq_zero, nat.cast_eq_zero] at h - -theorem re_eq_add_conj (z : K) : ↑(re z) = (z + conj z) / 2 := -by rw [add_conj, mul_div_cancel_left ((re z):K) two_ne_zero] - -theorem im_eq_conj_sub (z : K) : ↑(im z) = I * (conj z - z) / 2 := -begin - rw [← neg_inj, ← of_real_neg, ← I_mul_re, re_eq_add_conj], - simp only [mul_add, sub_eq_add_neg, neg_div', neg_mul, conj_I, - mul_neg, neg_add_rev, neg_neg, ring_hom.map_mul] -end - -/-! ### Absolute value -/ - -/-- The complex absolute value function, defined as the square root of the norm squared. -/ -@[pp_nodot] noncomputable def abs (z : K) : ℝ := (norm_sq z).sqrt - -local notation `abs'` := has_abs.abs -local notation `absK` := @abs K _ - -@[simp, norm_cast] lemma abs_of_real (r : ℝ) : absK r = abs' r := -by simp only [abs, norm_sq, real.sqrt_mul_self_eq_abs, add_zero, of_real_im, - monoid_with_zero_hom.coe_mk, of_real_re, mul_zero] - -lemma norm_eq_abs (z : K) : ‖z‖ = absK z := -by simp only [abs, norm_sq_eq_def', norm_nonneg, real.sqrt_sq] - -@[is_R_or_C_simps, norm_cast] -lemma norm_of_real (z : ℝ) : ‖(z : K)‖ = ‖z‖ := -by { rw [is_R_or_C.norm_eq_abs, is_R_or_C.abs_of_real, real.norm_eq_abs] } - -lemma abs_of_nonneg {r : ℝ} (h : 0 ≤ r) : absK r = r := -(abs_of_real _).trans (abs_of_nonneg h) - -lemma norm_of_nonneg {r : ℝ} (r_nn : 0 ≤ r) : ‖(r : K)‖ = r := -by { rw norm_of_real, exact abs_eq_self.mpr r_nn, } +/-! ### Norm -/ -lemma abs_of_nat (n : ℕ) : absK n = n := -by { rw [← of_real_nat_cast], exact abs_of_nonneg (nat.cast_nonneg n) } +lemma norm_of_nonneg {r : ℝ} (h : 0 ≤ r) : ‖(r : K)‖ = r := +(norm_of_real _).trans (abs_of_nonneg h) -lemma mul_self_abs (z : K) : abs z * abs z = norm_sq z := -real.mul_self_sqrt (norm_sq_nonneg _) +@[simp, priority 900, is_R_or_C_simps, norm_cast] +lemma norm_nat_cast (n : ℕ) : ‖(n : K)‖ = n := +by { rw [← of_real_nat_cast], exact norm_of_nonneg (nat.cast_nonneg n) } -@[simp, is_R_or_C_simps] lemma abs_zero : absK 0 = 0 := by simp only [abs, real.sqrt_zero, map_zero] -@[simp, is_R_or_C_simps] lemma abs_one : absK 1 = 1 := by simp only [abs, map_one, real.sqrt_one] +lemma mul_self_norm (z : K) : ‖z‖ * ‖z‖ = norm_sq z := +by rw [norm_sq_eq_def', sq] -@[simp, is_R_or_C_simps] lemma abs_two : absK 2 = 2 := -calc absK 2 = absK (2 : ℝ) : by rw [of_real_bit0, of_real_one] -... = (2 : ℝ) : abs_of_nonneg (by norm_num) +attribute [is_R_or_C_simps] norm_zero norm_one norm_eq_zero abs_norm norm_inv norm_div -lemma abs_nonneg (z : K) : 0 ≤ absK z := -real.sqrt_nonneg _ +@[simp, priority 900, is_R_or_C_simps] lemma norm_two : ‖(2 : K)‖ = 2 := +by rw [← nat.cast_two, norm_nat_cast, nat.cast_two] -@[simp, is_R_or_C_simps] lemma abs_eq_zero {z : K} : absK z = 0 ↔ z = 0 := -(real.sqrt_eq_zero $ norm_sq_nonneg _).trans norm_sq_eq_zero - -lemma abs_ne_zero {z : K} : abs z ≠ 0 ↔ z ≠ 0 := -not_congr abs_eq_zero - -@[simp, is_R_or_C_simps] lemma abs_conj (z : K) : abs (conj z) = abs z := -by simp only [abs, norm_sq_conj] - -@[simp, is_R_or_C_simps] lemma abs_mul (z w : K) : abs (z * w) = abs z * abs w := -by rw [abs, norm_sq_mul, real.sqrt_mul (norm_sq_nonneg _)]; refl - -lemma abs_re_le_abs (z : K) : abs' (re z) ≤ abs z := -by rw [mul_self_le_mul_self_iff (_root_.abs_nonneg (re z)) (abs_nonneg _), - abs_mul_abs_self, mul_self_abs]; +lemma abs_re_le_norm (z : K) : |re z| ≤ ‖z‖ := +by rw [mul_self_le_mul_self_iff (_root_.abs_nonneg (re z)) (norm_nonneg _), + abs_mul_abs_self, mul_self_norm]; apply re_sq_le_norm_sq -lemma abs_im_le_abs (z : K) : abs' (im z) ≤ abs z := -by rw [mul_self_le_mul_self_iff (_root_.abs_nonneg (im z)) (abs_nonneg _), - abs_mul_abs_self, mul_self_abs]; +lemma abs_im_le_norm (z : K) : |im z| ≤ ‖z‖ := +by rw [mul_self_le_mul_self_iff (_root_.abs_nonneg (im z)) (norm_nonneg _), + abs_mul_abs_self, mul_self_norm]; apply im_sq_le_norm_sq -lemma norm_re_le_norm (z : K) : ‖re z‖ ≤ ‖z‖ := -by { rw [is_R_or_C.norm_eq_abs, real.norm_eq_abs], exact is_R_or_C.abs_re_le_abs _, } - -lemma norm_im_le_norm (z : K) : ‖im z‖ ≤ ‖z‖ := -by { rw [is_R_or_C.norm_eq_abs, real.norm_eq_abs], exact is_R_or_C.abs_im_le_abs _, } +lemma norm_re_le_norm (z : K) : ‖re z‖ ≤ ‖z‖ := abs_re_le_norm z +lemma norm_im_le_norm (z : K) : ‖im z‖ ≤ ‖z‖ := abs_im_le_norm z -lemma re_le_abs (z : K) : re z ≤ abs z := -(abs_le.1 (abs_re_le_abs _)).2 +lemma re_le_norm (z : K) : re z ≤ ‖z‖ := (abs_le.1 (abs_re_le_norm z)).2 +lemma im_le_norm (z : K) : im z ≤ ‖z‖ := (abs_le.1 (abs_im_le_norm _)).2 -lemma im_le_abs (z : K) : im z ≤ abs z := -(abs_le.1 (abs_im_le_abs _)).2 +lemma im_eq_zero_of_le {a : K} (h : ‖a‖ ≤ re a) : im a = 0 := +by simpa only [mul_self_norm a, norm_sq_apply, self_eq_add_right, mul_self_eq_zero] + using congr_arg (λ z, z * z) ((re_le_norm a).antisymm h) -lemma im_eq_zero_of_le {a : K} (h : abs a ≤ re a) : im a = 0 := -begin - rw ← zero_eq_mul_self, - have : re a * re a = re a * re a + im a * im a, - { convert is_R_or_C.mul_self_abs a; - linarith [re_le_abs a] }, - linarith -end - -lemma re_eq_self_of_le {a : K} (h : abs a ≤ re a) : (re a : K) = a := -by { rw ← re_add_im a, simp only [im_eq_zero_of_le h, add_zero, zero_mul, algebra_map.coe_zero] - with is_R_or_C_simps, } +lemma re_eq_self_of_le {a : K} (h : ‖a‖ ≤ re a) : (re a : K) = a := +by rw [(is_real_tfae a).out 2 3, im_eq_zero_of_le h] -lemma abs_add (z w : K) : abs (z + w) ≤ abs z + abs w := -(mul_self_le_mul_self_iff (abs_nonneg _) - (add_nonneg (abs_nonneg _) (abs_nonneg _))).2 $ -begin - rw [mul_self_abs, add_mul_self_eq, mul_self_abs, mul_self_abs, - add_right_comm, norm_sq_add, add_le_add_iff_left, - mul_assoc, mul_le_mul_left (zero_lt_two' ℝ)], - simpa [-mul_re] with is_R_or_C_simps using re_le_abs (z * conj w) -end - -instance : is_absolute_value absK := -{ abv_nonneg := abs_nonneg, - abv_eq_zero := λ _, abs_eq_zero, - abv_add := abs_add, - abv_mul := abs_mul } open is_absolute_value -@[simp, is_R_or_C_simps] lemma abs_abs (z : K) : abs' (abs z) = abs z := -_root_.abs_of_nonneg (abs_nonneg _) - -@[simp, is_R_or_C_simps] lemma abs_pos {z : K} : 0 < abs z ↔ z ≠ 0 := abv_pos abs -@[simp, is_R_or_C_simps] lemma abs_neg : ∀ z : K, abs (-z) = abs z := abv_neg abs -lemma abs_sub : ∀ z w : K, abs (z - w) = abs (w - z) := abv_sub abs -lemma abs_sub_le : ∀ a b c : K, abs (a - c) ≤ abs (a - b) + abs (b - c) := abv_sub_le abs -@[simp, is_R_or_C_simps] theorem abs_inv : ∀ z : K, abs z⁻¹ = (abs z)⁻¹ := abv_inv abs -@[simp, is_R_or_C_simps] theorem abs_div : ∀ z w : K, abs (z / w) = abs z / abs w := abv_div abs - -lemma abs_abs_sub_le_abs_sub : ∀ z w : K, abs' (abs z - abs w) ≤ abs (z - w) := -abs_abv_sub_le_abv_sub abs - -lemma abs_re_div_abs_le_one (z : K) : abs' (re z / abs z) ≤ 1 := +lemma abs_re_div_norm_le_one (z : K) : |re z / ‖z‖| ≤ 1 := begin - by_cases hz : z = 0, - { simp [hz, zero_le_one] }, - { simp_rw [_root_.abs_div, abs_abs, div_le_iff (abs_pos.2 hz), one_mul, abs_re_le_abs] } + rw [abs_div, abs_norm], + exact div_le_one_of_le (abs_re_le_norm _) (norm_nonneg _) end -lemma abs_im_div_abs_le_one (z : K) : abs' (im z / abs z) ≤ 1 := +lemma abs_im_div_norm_le_one (z : K) : |im z / ‖z‖| ≤ 1 := begin - by_cases hz : z = 0, - { simp [hz, zero_le_one] }, - { simp_rw [_root_.abs_div, abs_abs, div_le_iff (abs_pos.2 hz), one_mul, abs_im_le_abs] } + rw [abs_div, abs_norm], + exact div_le_one_of_le (abs_im_le_norm _) (norm_nonneg _) end -@[simp, is_R_or_C_simps, norm_cast] lemma abs_cast_nat (n : ℕ) : abs (n : K) = n := -by rw [← of_real_nat_cast, abs_of_nonneg (nat.cast_nonneg n)] - -lemma norm_sq_eq_abs (x : K) : norm_sq x = abs x ^ 2 := -by rw [abs, sq, real.mul_self_sqrt (norm_sq_nonneg _)] +lemma re_eq_norm_of_mul_conj (x : K) : re (x * conj x) = ‖x * conj x‖ := +by rw [mul_conj, of_real_re, norm_of_real, abs_of_nonneg (norm_sq_nonneg _)] -lemma re_eq_abs_of_mul_conj (x : K) : re (x * (conj x)) = abs (x * (conj x)) := -by rw [mul_conj, of_real_re, abs_of_real, norm_sq_eq_abs, sq, _root_.abs_mul, abs_abs] +lemma norm_sq_re_add_conj (x : K) : (‖x + conj x‖)^2 = (re (x + conj x))^2 := +by rw [add_conj, norm_mul, norm_two, norm_of_real, two_mul (re x : K), map_add, of_real_re, + ← two_mul, mul_pow, mul_pow, sq_abs] -lemma abs_sq_re_add_conj (x : K) : (abs (x + conj x))^2 = (re (x + conj x))^2 := -by simp only [sq, ←norm_sq_eq_abs, norm_sq, map_add, add_zero, monoid_with_zero_hom.coe_mk, - add_right_neg, mul_zero] with is_R_or_C_simps - -lemma abs_sq_re_add_conj' (x : K) : (abs (conj x + x))^2 = (re (conj x + x))^2 := -by simp only [sq, ←norm_sq_eq_abs, norm_sq, map_add, add_zero, monoid_with_zero_hom.coe_mk, - add_left_neg, mul_zero] with is_R_or_C_simps +lemma norm_sq_re_conj_add (x : K) : (‖conj x + x‖)^2 = (re (conj x + x))^2 := +by rw [add_comm, norm_sq_re_add_conj] /-! ### Cauchy sequences -/ -theorem is_cau_seq_re (f : cau_seq K abs) : is_cau_seq abs' (λ n, re (f n)) := +theorem is_cau_seq_re (f : cau_seq K norm) : is_cau_seq abs (λ n, re (f n)) := λ ε ε0, (f.cauchy ε0).imp $ λ i H j ij, -lt_of_le_of_lt (by simpa using abs_re_le_abs (f j - f i)) (H _ ij) +lt_of_le_of_lt (by simpa only [map_sub] using abs_re_le_norm (f j - f i)) (H _ ij) -theorem is_cau_seq_im (f : cau_seq K abs) : is_cau_seq abs' (λ n, im (f n)) := +theorem is_cau_seq_im (f : cau_seq K norm) : is_cau_seq abs (λ n, im (f n)) := λ ε ε0, (f.cauchy ε0).imp $ λ i H j ij, -lt_of_le_of_lt (by simpa using abs_im_le_abs (f j - f i)) (H _ ij) +lt_of_le_of_lt (by simpa only [map_sub] using abs_im_le_norm (f j - f i)) (H _ ij) /-- The real part of a K Cauchy sequence, as a real Cauchy sequence. -/ -noncomputable def cau_seq_re (f : cau_seq K abs) : cau_seq ℝ abs' := +noncomputable def cau_seq_re (f : cau_seq K norm) : cau_seq ℝ abs := ⟨_, is_cau_seq_re f⟩ /-- The imaginary part of a K Cauchy sequence, as a real Cauchy sequence. -/ -noncomputable def cau_seq_im (f : cau_seq K abs) : cau_seq ℝ abs' := +noncomputable def cau_seq_im (f : cau_seq K norm) : cau_seq ℝ abs := ⟨_, is_cau_seq_im f⟩ -lemma is_cau_seq_abs {f : ℕ → K} (hf : is_cau_seq abs f) : - is_cau_seq abs' (abs ∘ f) := +lemma is_cau_seq_norm {f : ℕ → K} (hf : is_cau_seq norm f) : + is_cau_seq abs (norm ∘ f) := λ ε ε0, let ⟨i, hi⟩ := hf ε ε0 in -⟨i, λ j hj, lt_of_le_of_lt (abs_abs_sub_le_abs_sub _ _) (hi j hj)⟩ +⟨i, λ j hj, lt_of_le_of_lt (abs_norm_sub_norm_le _ _) (hi j hj)⟩ end is_R_or_C @@ -652,7 +560,6 @@ section cleanup_lemmas local notation `reR` := @is_R_or_C.re ℝ _ local notation `imR` := @is_R_or_C.im ℝ _ local notation `IR` := @is_R_or_C.I ℝ _ -local notation `absR` := @is_R_or_C.abs ℝ _ local notation `norm_sqR` := @is_R_or_C.norm_sq ℝ _ @[simp, is_R_or_C_simps] lemma re_to_real {x : ℝ} : reR x = x := rfl @@ -661,8 +568,6 @@ local notation `norm_sqR` := @is_R_or_C.norm_sq ℝ _ @[simp, is_R_or_C_simps] lemma I_to_real : IR = 0 := rfl @[simp, is_R_or_C_simps] lemma norm_sq_to_real {x : ℝ} : norm_sq x = x*x := by simp [is_R_or_C.norm_sq] -@[simp, is_R_or_C_simps] lemma abs_to_real {x : ℝ} : absR x = has_abs.abs x := -by simp [is_R_or_C.abs, abs, real.sqrt_mul_self_eq_abs] @[simp] lemma coe_real_eq_id : @coe ℝ ℝ _ = id := rfl @@ -678,8 +583,7 @@ def re_lm : K →ₗ[ℝ] ℝ := /-- The real part in a `is_R_or_C` field, as a continuous linear map. -/ noncomputable def re_clm : K →L[ℝ] ℝ := -linear_map.mk_continuous re_lm 1 $ by -{ simp only [norm_eq_abs, re_lm_coe, one_mul, abs_to_real], exact abs_re_le_abs, } +linear_map.mk_continuous re_lm 1 $ λ x, by { rw [one_mul], exact abs_re_le_norm x } @[simp, is_R_or_C_simps, norm_cast] lemma re_clm_coe : ((re_clm : K →L[ℝ] ℝ) : K →ₗ[ℝ] ℝ) = re_lm := rfl @@ -696,8 +600,7 @@ def im_lm : K →ₗ[ℝ] ℝ := /-- The imaginary part in a `is_R_or_C` field, as a continuous linear map. -/ noncomputable def im_clm : K →L[ℝ] ℝ := -linear_map.mk_continuous im_lm 1 $ by -{ simp only [norm_eq_abs, re_lm_coe, one_mul, abs_to_real], exact abs_im_le_abs, } +linear_map.mk_continuous im_lm 1 $ fun x, by { rw [one_mul], exact abs_im_le_norm x } @[simp, is_R_or_C_simps, norm_cast] lemma im_clm_coe : ((im_clm : K →L[ℝ] ℝ) : K →ₗ[ℝ] ℝ) = im_lm := rfl @@ -718,7 +621,7 @@ def conj_ae : K ≃ₐ[ℝ] K := /-- Conjugate as a linear isometry -/ noncomputable def conj_lie : K ≃ₗᵢ[ℝ] K := -⟨conj_ae.to_linear_equiv, λ z, by simp [norm_eq_abs] with is_R_or_C_simps⟩ +⟨conj_ae.to_linear_equiv, λ _, norm_conj⟩ @[simp, is_R_or_C_simps] lemma conj_lie_apply : (conj_lie : K → K) = conj := rfl @@ -742,7 +645,7 @@ noncomputable def of_real_am : ℝ →ₐ[ℝ] K := algebra.of_id ℝ K /-- The ℝ → K coercion, as a linear isometry -/ noncomputable def of_real_li : ℝ →ₗᵢ[ℝ] K := -{ to_linear_map := of_real_am.to_linear_map, norm_map' := by simp [norm_eq_abs] } +{ to_linear_map := of_real_am.to_linear_map, norm_map' := norm_of_real } @[simp, is_R_or_C_simps] lemma of_real_li_apply : (of_real_li : ℝ → K) = coe := rfl @@ -756,17 +659,8 @@ noncomputable def of_real_clm : ℝ →L[ℝ] K := of_real_li.to_continuous_line @[continuity] lemma continuous_of_real : continuous (coe : ℝ → K) := of_real_li.continuous -@[continuity] lemma continuous_abs : continuous (@is_R_or_C.abs K _) := -by simp only [show @is_R_or_C.abs K _ = has_norm.norm, by { ext, exact (norm_eq_abs _).symm }, - continuous_norm] - -@[continuity] lemma continuous_norm_sq : continuous (@is_R_or_C.norm_sq K _) := -begin - have : (@is_R_or_C.norm_sq K _ : K → ℝ) = λ x, (is_R_or_C.abs x) ^ 2, - { ext, - exact norm_sq_eq_abs _ }, - simp only [this, continuous_abs.pow 2], -end +@[continuity] lemma continuous_norm_sq : continuous (norm_sq : K → ℝ) := +(continuous_re.mul continuous_re).add (continuous_im.mul continuous_im) end linear_maps diff --git a/src/measure_theory/function/l2_space.lean b/src/measure_theory/function/l2_space.lean index b7b913670cbc8..ee1620be1b073 100644 --- a/src/measure_theory/function/l2_space.lean +++ b/src/measure_theory/function/l2_space.lean @@ -72,19 +72,17 @@ end lemma snorm_inner_lt_top (f g : α →₂[μ] E) : snorm (λ (x : α), ⟪f x, g x⟫) 1 μ < ∞ := begin - have h : ∀ x, is_R_or_C.abs ⟪f x, g x⟫ ≤ ‖f x‖ * ‖g x‖, from λ x, abs_inner_le_norm _ _, - have h' : ∀ x, is_R_or_C.abs ⟪f x, g x⟫ ≤ is_R_or_C.abs (‖f x‖^2 + ‖g x‖^2), - { refine λ x, le_trans (h x) _, - rw [is_R_or_C.abs_to_real, abs_eq_self.mpr], - swap, { exact add_nonneg (by simp) (by simp), }, - refine le_trans _ (half_le_self (add_nonneg (sq_nonneg _) (sq_nonneg _))), - refine (le_div_iff (zero_lt_two' ℝ)).mpr ((le_of_eq _).trans (two_mul_le_add_sq _ _)), - ring, }, - simp_rw [← is_R_or_C.norm_eq_abs, ← real.rpow_nat_cast] at h', - refine (snorm_mono_ae (ae_of_all _ h')).trans_lt ((snorm_add_le _ _ le_rfl).trans_lt _), + have h : ∀ x, ‖⟪f x, g x⟫‖ ≤ ‖‖f x‖ ^ (2 : ℝ) + ‖g x‖ ^ (2 : ℝ)‖, + { intro x, + rw [← @nat.cast_two ℝ, real.rpow_nat_cast, real.rpow_nat_cast], + calc ‖⟪f x, g x⟫‖ ≤ ‖f x‖ * ‖g x‖ : norm_inner_le_norm _ _ + ... ≤ 2 * ‖f x‖ * ‖g x‖ : + mul_le_mul_of_nonneg_right (le_mul_of_one_le_left (norm_nonneg _) one_le_two) (norm_nonneg _) + ... ≤ ‖‖f x‖^2 + ‖g x‖^2‖ : (two_mul_le_add_sq _ _).trans (le_abs_self _) }, + refine (snorm_mono_ae (ae_of_all _ h)).trans_lt ((snorm_add_le _ _ le_rfl).trans_lt _), { exact ((Lp.ae_strongly_measurable f).norm.ae_measurable.pow_const _).ae_strongly_measurable }, { exact ((Lp.ae_strongly_measurable g).norm.ae_measurable.pow_const _).ae_strongly_measurable }, - simp only [nat.cast_bit0, ennreal.add_lt_top, nat.cast_one], + rw [ennreal.add_lt_top], exact ⟨snorm_rpow_two_norm_lt_top f, snorm_rpow_two_norm_lt_top g⟩, end diff --git a/src/ring_theory/polynomial/cyclotomic/eval.lean b/src/ring_theory/polynomial/cyclotomic/eval.lean index 07d7e226483ec..77a8f285d4c51 100644 --- a/src/ring_theory/polynomial/cyclotomic/eval.lean +++ b/src/ring_theory/polynomial/cyclotomic/eval.lean @@ -253,7 +253,7 @@ begin { refine ⟨ζ, (mem_primitive_roots hn).mpr hζ, _⟩, suffices : ¬ same_ray ℝ (q : ℂ) (-ζ), { convert norm_add_lt_of_not_same_ray this; - simp [real.norm_of_nonneg hq.le, hζ.norm'_eq_one hn.ne', -complex.norm_eq_abs] }, + simp [abs_of_pos hq, hζ.norm'_eq_one hn.ne', -complex.norm_eq_abs] }, rw complex.same_ray_iff, push_neg, refine ⟨by exact_mod_cast hq.ne', neg_ne_zero.mpr $ hζ.ne_zero hn.ne', _⟩, From 937b1c59c58710ef8ed91f8727ef402d49d621a2 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 8 May 2023 06:19:52 +0000 Subject: [PATCH 0970/1291] feat(order/category): Category of finite distributive lattices (#11677) Define `FinBddDistLat`, the category of finite distributive lattices with bounded lattice homomorphisms. This is one of the two categories involved in Birkhoff's representation theorem. --- src/order/category/FinBddDistLat.lean | 83 +++++++++++++++++++++++++++ src/order/category/FinBoolAlg.lean | 10 +++- src/order/category/FinPartOrd.lean | 4 +- 3 files changed, 94 insertions(+), 3 deletions(-) create mode 100644 src/order/category/FinBddDistLat.lean diff --git a/src/order/category/FinBddDistLat.lean b/src/order/category/FinBddDistLat.lean new file mode 100644 index 0000000000000..2e04a3deab863 --- /dev/null +++ b/src/order/category/FinBddDistLat.lean @@ -0,0 +1,83 @@ +/- +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.fintype.order +import order.category.BddDistLat +import order.category.FinPartOrd + +/-! +# The category of finite bounded distributive lattices + +This file defines `FinBddDistLat`, the category of finite distributive lattices with +bounded lattice homomorphisms. +-/ + +universes u + +open category_theory + +/-- The category of finite distributive lattices with bounded lattice morphisms. -/ +structure FinBddDistLat := +(to_BddDistLat : BddDistLat) +[is_fintype : fintype to_BddDistLat] + +namespace FinBddDistLat + +instance : has_coe_to_sort FinBddDistLat Type* := ⟨λ X, X.to_BddDistLat⟩ +instance (X : FinBddDistLat) : distrib_lattice X := +X.to_BddDistLat.to_DistLat.str +instance (X : FinBddDistLat) : bounded_order X := X.to_BddDistLat.is_bounded_order + +attribute [instance] FinBddDistLat.is_fintype + +/-- Construct a bundled `FinBddDistLat` from a `nonempty` `bounded_order` `distrib_lattice`. -/ +def of (α : Type*) [distrib_lattice α] [bounded_order α] [fintype α] : FinBddDistLat := +⟨⟨⟨α⟩⟩⟩ + +/-- Construct a bundled `FinBddDistLat` from a `nonempty` `bounded_order` `distrib_lattice`. -/ +def of' (α : Type*) [distrib_lattice α] [fintype α] [nonempty α] : FinBddDistLat := +by { haveI := fintype.to_bounded_order α, exact ⟨⟨⟨α⟩⟩⟩ } + +instance : inhabited FinBddDistLat := ⟨of punit⟩ + +instance large_category : large_category FinBddDistLat := +induced_category.category to_BddDistLat + +instance concrete_category : concrete_category FinBddDistLat := +induced_category.concrete_category to_BddDistLat + +instance has_forget_to_BddDistLat : + has_forget₂ FinBddDistLat BddDistLat := +induced_category.has_forget₂ FinBddDistLat.to_BddDistLat + +instance has_forget_to_FinPartOrd : has_forget₂ FinBddDistLat FinPartOrd := +{ forget₂ := { obj := λ X, FinPartOrd.of X, + map := λ X Y f, (show bounded_lattice_hom X Y, from f : X →o Y) } } + +/-- Constructs an equivalence between finite distributive lattices from an order isomorphism +between them. -/ +@[simps] def iso.mk {α β : FinBddDistLat.{u}} (e : α ≃o β) : α ≅ β := +{ hom := (e : bounded_lattice_hom α β), + inv := (e.symm : bounded_lattice_hom β α), + hom_inv_id' := by { ext, exact e.symm_apply_apply _ }, + inv_hom_id' := by { ext, exact e.apply_symm_apply _ } } + +example {X Y : FinBddDistLat} : (X ⟶ Y) = bounded_lattice_hom X Y := rfl + +/-- `order_dual` as a functor. -/ +@[simps] def dual : FinBddDistLat ⥤ FinBddDistLat := +{ obj := λ X, of Xᵒᵈ, map := λ X Y, bounded_lattice_hom.dual } + +/-- The equivalence between `FinBddDistLat` and itself induced by `order_dual` both ways. -/ +@[simps functor inverse] def dual_equiv : FinBddDistLat ≌ FinBddDistLat := +equivalence.mk dual dual + (nat_iso.of_components (λ X, iso.mk $ order_iso.dual_dual X) $ λ X Y f, rfl) + (nat_iso.of_components (λ X, iso.mk $ order_iso.dual_dual X) $ λ X Y f, rfl) + +end FinBddDistLat + +lemma FinBddDistLat_dual_comp_forget_to_BddDistLat : + FinBddDistLat.dual ⋙ forget₂ FinBddDistLat BddDistLat = + forget₂ FinBddDistLat BddDistLat ⋙ BddDistLat.dual := rfl diff --git a/src/order/category/FinBoolAlg.lean b/src/order/category/FinBoolAlg.lean index 4da81128cd7cd..5082fe335a6ed 100644 --- a/src/order/category/FinBoolAlg.lean +++ b/src/order/category/FinBoolAlg.lean @@ -5,7 +5,7 @@ Authors: Yaël Dillies -/ import data.fintype.powerset import order.category.BoolAlg -import order.category.FinPartOrd +import order.category.FinBddDistLat import order.hom.complete_lattice /-! @@ -56,6 +56,10 @@ induced_category.concrete_category FinBoolAlg.to_BoolAlg instance has_forget_to_BoolAlg : has_forget₂ FinBoolAlg BoolAlg := induced_category.has_forget₂ FinBoolAlg.to_BoolAlg +instance has_forget_to_FinBddDistLat : has_forget₂ FinBoolAlg FinBddDistLat := +{ forget₂ := { obj := λ X, FinBddDistLat.of X, map := λ X Y f, f }, + forget_comp := rfl } + instance forget_to_BoolAlg_full : full (forget₂ FinBoolAlg BoolAlg) := induced_category.full _ instance forget_to_BoolAlg_faithful : faithful (forget₂ FinBoolAlg BoolAlg) := induced_category.faithful _ @@ -87,6 +91,10 @@ equivalence.mk dual dual end FinBoolAlg +lemma FinBoolAlg_dual_comp_forget_to_FinBddDistLat : + FinBoolAlg.dual ⋙ forget₂ FinBoolAlg FinBddDistLat = + forget₂ FinBoolAlg FinBddDistLat ⋙ FinBddDistLat.dual := rfl + /-- The powerset functor. `set` as a functor. -/ @[simps] def Fintype_to_FinBoolAlg_op : Fintype ⥤ FinBoolAlgᵒᵖ := { obj := λ X, op $ FinBoolAlg.of (set X), diff --git a/src/order/category/FinPartOrd.lean b/src/order/category/FinPartOrd.lean index 189d468eded93..ffa0dab0797f1 100644 --- a/src/order/category/FinPartOrd.lean +++ b/src/order/category/FinPartOrd.lean @@ -11,8 +11,8 @@ import order.category.PartOrd This defines `FinPartOrd`, the category of finite partial orders. -Note: `FinPartOrd` is NOT a subcategory of `BddOrd` because its morphisms do not -preserve `⊥` and `⊤`. +Note: `FinPartOrd` is *not* a subcategory of `BddOrd` because finite orders are not necessarily +bounded. ## TODO From f81174bd8aa3896304e8e1c1cb1d11cb0e0a05df Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 8 May 2023 07:29:47 +0000 Subject: [PATCH 0971/1291] feat(analysis/calculus/deriv): `polynomial.aeval` lemmas (#18945) This duplicates every lemma about differentiation of `polynomial.eval` for `polynomial.aeval` too. Some of these turned out to be needed in #18896. --- src/analysis/calculus/deriv.lean | 54 ++++++++++++++++++++++++++++++-- 1 file changed, 52 insertions(+), 2 deletions(-) diff --git a/src/analysis/calculus/deriv.lean b/src/analysis/calculus/deriv.lean index d3f5f5e96cce2..20327a2f4cc7b 100644 --- a/src/analysis/calculus/deriv.lean +++ b/src/analysis/calculus/deriv.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Gabriel Ebner, Sébastien Gouëzel -/ import analysis.calculus.fderiv +import data.polynomial.algebra_map import data.polynomial.derivative import linear_algebra.affine_space.slope @@ -1837,8 +1838,8 @@ end namespace polynomial /-! ### Derivative of a polynomial -/ -variables {x : 𝕜} {s : set 𝕜} -variable (p : 𝕜[X]) +variables {R : Type*} [comm_semiring R] [algebra R 𝕜] +variables (p : 𝕜[X]) (q : R[X]) {x : 𝕜} {s : set 𝕜} /-- The derivative (in the analysis sense) of a polynomial `p` is given by `p.derivative`. -/ protected lemma has_strict_deriv_at (x : 𝕜) : @@ -1857,29 +1858,57 @@ begin id.def], ring } } end +protected lemma has_strict_deriv_at_aeval (x : 𝕜) : + has_strict_deriv_at (λx, aeval x q) (aeval x q.derivative) x := +by simpa only [aeval_def, eval₂_eq_eval_map, derivative_map] + using (q.map (algebra_map R 𝕜)).has_strict_deriv_at x + /-- The derivative (in the analysis sense) of a polynomial `p` is given by `p.derivative`. -/ protected lemma has_deriv_at (x : 𝕜) : has_deriv_at (λx, p.eval x) (p.derivative.eval x) x := (p.has_strict_deriv_at x).has_deriv_at +protected lemma has_deriv_at_aeval (x : 𝕜) : + has_deriv_at (λx, aeval x q) (aeval x q.derivative) x := +(q.has_strict_deriv_at_aeval x).has_deriv_at + protected theorem has_deriv_within_at (x : 𝕜) (s : set 𝕜) : has_deriv_within_at (λx, p.eval x) (p.derivative.eval x) s x := (p.has_deriv_at x).has_deriv_within_at +protected theorem has_deriv_within_at_aeval (x : 𝕜) (s : set 𝕜) : + has_deriv_within_at (λx, aeval x q) (aeval x q.derivative) s x := +(q.has_deriv_at_aeval x).has_deriv_within_at + protected lemma differentiable_at : differentiable_at 𝕜 (λx, p.eval x) x := (p.has_deriv_at x).differentiable_at +protected lemma differentiable_at_aeval : differentiable_at 𝕜 (λx, aeval x q) x := +(q.has_deriv_at_aeval x).differentiable_at + protected lemma differentiable_within_at : differentiable_within_at 𝕜 (λx, p.eval x) s x := p.differentiable_at.differentiable_within_at +protected lemma differentiable_within_at_aeval : differentiable_within_at 𝕜 (λx, aeval x q) s x := +q.differentiable_at_aeval.differentiable_within_at + protected lemma differentiable : differentiable 𝕜 (λx, p.eval x) := λx, p.differentiable_at +protected lemma differentiable_aeval : differentiable 𝕜 (λ x : 𝕜, aeval x q) := +λx, q.differentiable_at_aeval + protected lemma differentiable_on : differentiable_on 𝕜 (λx, p.eval x) s := p.differentiable.differentiable_on +protected lemma differentiable_on_aeval : differentiable_on 𝕜 (λx, aeval x q) s := +q.differentiable_aeval.differentiable_on + @[simp] protected lemma deriv : deriv (λx, p.eval x) x = p.derivative.eval x := (p.has_deriv_at x).deriv +@[simp] protected lemma deriv_aeval : deriv (λx, aeval x q) x = aeval x q.derivative := +(q.has_deriv_at_aeval x).deriv + protected lemma deriv_within (hxs : unique_diff_within_at 𝕜 s x) : deriv_within (λx, p.eval x) s x = p.derivative.eval x := begin @@ -1887,22 +1916,43 @@ begin exact p.deriv end +protected lemma deriv_within_aeval (hxs : unique_diff_within_at 𝕜 s x) : + deriv_within (λx, aeval x q) s x = aeval x q.derivative := +by simpa only [aeval_def, eval₂_eq_eval_map, derivative_map] + using (q.map (algebra_map R 𝕜)).deriv_within hxs + protected lemma has_fderiv_at (x : 𝕜) : has_fderiv_at (λx, p.eval x) (smul_right (1 : 𝕜 →L[𝕜] 𝕜) (p.derivative.eval x)) x := p.has_deriv_at x +protected lemma has_fderiv_at_aeval (x : 𝕜) : + has_fderiv_at (λx, aeval x q) (smul_right (1 : 𝕜 →L[𝕜] 𝕜) (aeval x q.derivative)) x := +q.has_deriv_at_aeval x + protected lemma has_fderiv_within_at (x : 𝕜) : has_fderiv_within_at (λx, p.eval x) (smul_right (1 : 𝕜 →L[𝕜] 𝕜) (p.derivative.eval x)) s x := (p.has_fderiv_at x).has_fderiv_within_at +protected lemma has_fderiv_within_at_aeval (x : 𝕜) : + has_fderiv_within_at (λx, aeval x q) (smul_right (1 : 𝕜 →L[𝕜] 𝕜) (aeval x q.derivative)) s x := +(q.has_fderiv_at_aeval x).has_fderiv_within_at + @[simp] protected lemma fderiv : fderiv 𝕜 (λx, p.eval x) x = smul_right (1 : 𝕜 →L[𝕜] 𝕜) (p.derivative.eval x) := (p.has_fderiv_at x).fderiv +@[simp] protected lemma fderiv_aeval : + fderiv 𝕜 (λx, aeval x q) x = smul_right (1 : 𝕜 →L[𝕜] 𝕜) (aeval x q.derivative) := +(q.has_fderiv_at_aeval x).fderiv + protected lemma fderiv_within (hxs : unique_diff_within_at 𝕜 s x) : fderiv_within 𝕜 (λx, p.eval x) s x = smul_right (1 : 𝕜 →L[𝕜] 𝕜) (p.derivative.eval x) := (p.has_fderiv_within_at x).fderiv_within hxs +protected lemma fderiv_within_aeval (hxs : unique_diff_within_at 𝕜 s x) : + fderiv_within 𝕜 (λx, aeval x q) s x = smul_right (1 : 𝕜 →L[𝕜] 𝕜) (aeval x q.derivative) := +(q.has_fderiv_within_at_aeval x).fderiv_within hxs + end polynomial section pow From 6632ca2081e55ff5cf228ca63011979a0efb495b Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Mon, 8 May 2023 07:29:48 +0000 Subject: [PATCH 0972/1291] feat(algebra/order/group/order_iso): add `order_iso.div_{left,right}` (#18968) --- src/algebra/order/group/order_iso.lean | 12 ++++++++++++ 1 file changed, 12 insertions(+) diff --git a/src/algebra/order/group/order_iso.lean b/src/algebra/order/group/order_iso.lean index eadcf1a5b45a7..c82dca13aad5e 100644 --- a/src/algebra/order/group/order_iso.lean +++ b/src/algebra/order/group/order_iso.lean @@ -51,6 +51,12 @@ attribute [to_additive neg_le_of_neg_le] inv_le_of_inv_le' lemma le_inv' : a ≤ b⁻¹ ↔ b ≤ a⁻¹ := (order_iso.inv α).le_symm_apply +/-- `x ↦ a / x` as an order-reversing equivalence. -/ +@[to_additive "`x ↦ a - x` as an order-reversing equivalence.", simps] +def order_iso.div_left (a : α) : α ≃o αᵒᵈ := +{ to_equiv := (equiv.div_left a).trans order_dual.to_dual, + map_rel_iff' := λ x y, @div_le_div_iff_left α _ _ _ _ _ _ _ } + end typeclasses_left_right_le end group @@ -74,6 +80,12 @@ def order_iso.mul_right (a : α) : α ≃o α := (order_iso.mul_right a).symm = order_iso.mul_right a⁻¹ := by { ext x, refl } +/-- `x ↦ x / a` as an order isomorphism. -/ +@[to_additive "`x ↦ x - a` as an order isomorphism.", simps] +def order_iso.div_right (a : α) : α ≃o α := +{ to_equiv := equiv.div_right a, + map_rel_iff' := λ x y, div_le_div_iff_right a } + end right section left From 63decdb840b6213fdd846922a7becd822483d8b9 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Mon, 8 May 2023 10:37:08 +0000 Subject: [PATCH 0973/1291] feat(probability/kernel/basic): add `kernel.comap_right` and `kernel.piecewise` (#18917) Also put `is_s_finite_kernel` in the same namespace as the other classes of kernels and add `iff` variants of the ext lemmas. --- src/probability/kernel/basic.lean | 151 +++++++++++++++++++++++++++++- 1 file changed, 147 insertions(+), 4 deletions(-) diff --git a/src/probability/kernel/basic.lean b/src/probability/kernel/basic.lean index e4cadd40d6d8d..9a85dfb185e3f 100644 --- a/src/probability/kernel/basic.lean +++ b/src/probability/kernel/basic.lean @@ -26,7 +26,7 @@ Classes of kernels: particular that all measures in the image of `κ` are finite, but is stronger since it requires an uniform bound. This stronger condition is necessary to ensure that the composition of two finite kernels is finite. -* `probability_theory.kernel.is_s_finite_kernel κ`: a kernel is called s-finite if it is a countable +* `probability_theory.is_s_finite_kernel κ`: a kernel is called s-finite if it is a countable sum of finite kernels. Particular kernels: @@ -154,10 +154,16 @@ instance is_markov_kernel.is_finite_kernel [h : is_markov_kernel κ] : is_finite namespace kernel -@[ext] lemma ext {κ : kernel α β} {η : kernel α β} (h : ∀ a, κ a = η a) : κ = η := +@[ext] lemma ext {η : kernel α β} (h : ∀ a, κ a = η a) : κ = η := by { ext1, ext1 a, exact h a, } -lemma ext_fun {κ η : kernel α β} (h : ∀ a f, measurable f → ∫⁻ b, f b ∂(κ a) = ∫⁻ b, f b ∂(η a)) : +lemma ext_iff {η : kernel α β} : κ = η ↔ ∀ a, κ a = η a := +⟨λ h a, by rw h, ext⟩ + +lemma ext_iff' {η : kernel α β} : κ = η ↔ ∀ a (s : set β) (hs : measurable_set s), κ a s = η a s := +by simp_rw [ext_iff, measure.ext_iff] + +lemma ext_fun {η : kernel α β} (h : ∀ a f, measurable f → ∫⁻ b, f b ∂(κ a) = ∫⁻ b, f b ∂(η a)) : κ = η := begin ext a s hs, @@ -166,6 +172,10 @@ begin rw h, end +lemma ext_fun_iff {η : kernel α β} : + κ = η ↔ ∀ a f, measurable f → ∫⁻ b, f b ∂(κ a) = ∫⁻ b, f b ∂(η a) := +⟨λ h a f hf, by rw h, ext_fun⟩ + protected lemma measurable (κ : kernel α β) : measurable κ := κ.prop protected lemma measurable_coe (κ : kernel α β) {s : set β} (hs : measurable_set s) : @@ -220,7 +230,7 @@ end sum section s_finite /-- A kernel is s-finite if it can be written as the sum of countably many finite kernels. -/ -class is_s_finite_kernel (κ : kernel α β) : Prop := +class _root_.probability_theory.is_s_finite_kernel (κ : kernel α β) : Prop := (tsum_finite : ∃ κs : ℕ → kernel α β, (∀ n, is_finite_kernel (κs n)) ∧ κ = kernel.sum κs) @[priority 100] @@ -421,6 +431,139 @@ end end restrict +section comap_right + +variables {γ : Type*} {mγ : measurable_space γ} {f : γ → β} + +include mγ + +/-- Kernel with value `(κ a).comap f`, for a measurable embedding `f`. That is, for a measurable set +`t : set β`, `comap_right κ hf a t = κ a (f '' t)`. -/ +noncomputable +def comap_right (κ : kernel α β) (hf : measurable_embedding f) : + kernel α γ := +{ val := λ a, (κ a).comap f, + property := + begin + refine measure.measurable_measure.mpr (λ t ht, _), + have : (λ a, measure.comap f (κ a) t) = λ a, κ a (f '' t), + { ext1 a, + rw measure.comap_apply _ hf.injective (λ s' hs', _) _ ht, + exact hf.measurable_set_image.mpr hs', }, + rw this, + exact kernel.measurable_coe _ (hf.measurable_set_image.mpr ht), + end } + +lemma comap_right_apply (κ : kernel α β) (hf : measurable_embedding f) (a : α) : + comap_right κ hf a = measure.comap f (κ a) := rfl + +lemma comap_right_apply' (κ : kernel α β) (hf : measurable_embedding f) + (a : α) {t : set γ} (ht : measurable_set t) : + comap_right κ hf a t = κ a (f '' t) := +by rw [comap_right_apply, + measure.comap_apply _ hf.injective (λ s, hf.measurable_set_image.mpr) _ ht] + +lemma is_markov_kernel.comap_right (κ : kernel α β) (hf : measurable_embedding f) + (hκ : ∀ a, κ a (set.range f) = 1) : + is_markov_kernel (comap_right κ hf) := +begin + refine ⟨λ a, ⟨_⟩⟩, + rw comap_right_apply' κ hf a measurable_set.univ, + simp only [set.image_univ, subtype.range_coe_subtype, set.set_of_mem_eq], + exact hκ a, +end + +instance is_finite_kernel.comap_right (κ : kernel α β) [is_finite_kernel κ] + (hf : measurable_embedding f) : + is_finite_kernel (comap_right κ hf) := +begin + refine ⟨⟨is_finite_kernel.bound κ, is_finite_kernel.bound_lt_top κ, λ a, _⟩⟩, + rw comap_right_apply' κ hf a measurable_set.univ, + exact measure_le_bound κ a _, +end + +instance is_s_finite_kernel.comap_right (κ : kernel α β) [is_s_finite_kernel κ] + (hf : measurable_embedding f) : + is_s_finite_kernel (comap_right κ hf) := +begin + refine ⟨⟨λ n, comap_right (seq κ n) hf, infer_instance, _⟩⟩, + ext1 a, + rw sum_apply, + simp_rw comap_right_apply _ hf, + have : measure.sum (λ n, measure.comap f (seq κ n a)) + = measure.comap f (measure.sum (λ n, seq κ n a)), + { ext1 t ht, + rw [measure.comap_apply _ hf.injective (λ s', hf.measurable_set_image.mpr) _ ht, + measure.sum_apply _ ht, measure.sum_apply _ (hf.measurable_set_image.mpr ht)], + congr' with n : 1, + rw measure.comap_apply _ hf.injective (λ s', hf.measurable_set_image.mpr) _ ht, }, + rw [this, measure_sum_seq], +end + +end comap_right + +section piecewise + +variables {η : kernel α β} {s : set α} {hs : measurable_set s} [decidable_pred (∈ s)] + +/-- `piecewise hs κ η` is the kernel equal to `κ` on the measurable set `s` and to `η` on its +complement. -/ +def piecewise (hs : measurable_set s) (κ η : kernel α β) : + kernel α β := +{ val := λ a, if a ∈ s then κ a else η a, + property := measurable.piecewise hs (kernel.measurable _) (kernel.measurable _) } + +lemma piecewise_apply (a : α) : + piecewise hs κ η a = if a ∈ s then κ a else η a := rfl + +lemma piecewise_apply' (a : α) (t : set β) : + piecewise hs κ η a t = if a ∈ s then κ a t else η a t := +by { rw piecewise_apply, split_ifs; refl, } + +instance is_markov_kernel.piecewise [is_markov_kernel κ] [is_markov_kernel η] : + is_markov_kernel (piecewise hs κ η) := +by { refine ⟨λ a, ⟨_⟩⟩, rw [piecewise_apply', measure_univ, measure_univ, if_t_t], } + +instance is_finite_kernel.piecewise [is_finite_kernel κ] [is_finite_kernel η] : + is_finite_kernel (piecewise hs κ η) := +begin + refine ⟨⟨max (is_finite_kernel.bound κ) (is_finite_kernel.bound η), _, λ a, _⟩⟩, + { exact max_lt (is_finite_kernel.bound_lt_top κ) (is_finite_kernel.bound_lt_top η), }, + rw [piecewise_apply'], + exact (ite_le_sup _ _ _).trans (sup_le_sup (measure_le_bound _ _ _) (measure_le_bound _ _ _)), +end + +instance is_s_finite_kernel.piecewise [is_s_finite_kernel κ] [is_s_finite_kernel η] : + is_s_finite_kernel (piecewise hs κ η) := +begin + refine ⟨⟨λ n, piecewise hs (seq κ n) (seq η n), infer_instance, _⟩⟩, + ext1 a, + simp_rw [sum_apply, kernel.piecewise_apply], + split_ifs; exact (measure_sum_seq _ a).symm, +end + +lemma lintegral_piecewise (a : α) (g : β → ℝ≥0∞) : + ∫⁻ b, g b ∂(piecewise hs κ η a) = if a ∈ s then ∫⁻ b, g b ∂(κ a) else ∫⁻ b, g b ∂(η a) := +by { simp_rw piecewise_apply, split_ifs; refl, } + +lemma set_lintegral_piecewise (a : α) (g : β → ℝ≥0∞) (t : set β) : + ∫⁻ b in t, g b ∂(piecewise hs κ η a) + = if a ∈ s then ∫⁻ b in t, g b ∂(κ a) else ∫⁻ b in t, g b ∂(η a) := +by { simp_rw piecewise_apply, split_ifs; refl, } + +lemma integral_piecewise {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] [complete_space E] + (a : α) (g : β → E) : + ∫ b, g b ∂(piecewise hs κ η a) = if a ∈ s then ∫ b, g b ∂(κ a) else ∫ b, g b ∂(η a) := +by { simp_rw piecewise_apply, split_ifs; refl, } + +lemma set_integral_piecewise {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] + [complete_space E] (a : α) (g : β → E) (t : set β) : + ∫ b in t, g b ∂(piecewise hs κ η a) + = if a ∈ s then ∫ b in t, g b ∂(κ a) else ∫ b in t, g b ∂(η a) := +by { simp_rw piecewise_apply, split_ifs; refl, } + +end piecewise + end kernel end probability_theory From a9545e8a564bac7f24637443f52ae955474e4991 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Mon, 8 May 2023 10:37:09 +0000 Subject: [PATCH 0974/1291] feat(probability/kernel/basic): integrals of basic kernels (#18961) Lebesgue and Bochner integral of a function against deterministic, constant and restricted kernels. --- src/measure_theory/integral/bochner.lean | 21 ++++++ src/measure_theory/integral/lebesgue.lean | 61 ++++++++++++++++ src/probability/kernel/basic.lean | 89 +++++++++++++++++++++++ src/probability/kernel/composition.lean | 4 +- src/probability/kernel/invariance.lean | 2 +- 5 files changed, 174 insertions(+), 3 deletions(-) diff --git a/src/measure_theory/integral/bochner.lean b/src/measure_theory/integral/bochner.lean index a50d1b9eb1dfd..82ee783fd13b7 100644 --- a/src/measure_theory/integral/bochner.lean +++ b/src/measure_theory/integral/bochner.lean @@ -1532,6 +1532,27 @@ calc ∫ x, f x ∂(measure.dirac a) = ∫ x, f a ∂(measure.dirac a) : integral_congr_ae $ ae_eq_dirac f ... = f a : by simp [measure.dirac_apply_of_mem] +lemma set_integral_dirac' {mα : measurable_space α} {f : α → E} (hf : strongly_measurable f) + (a : α) {s : set α} (hs : measurable_set s) [decidable (a ∈ s)] : + ∫ x in s, f x ∂(measure.dirac a) = if a ∈ s then f a else 0 := +begin + rw [restrict_dirac' hs], + swap, { apply_instance, }, + split_ifs, + { exact integral_dirac' _ _ hf, }, + { exact integral_zero_measure _, }, +end + +lemma set_integral_dirac [measurable_space α] [measurable_singleton_class α] (f : α → E) + (a : α) (s : set α) [decidable (a ∈ s)] : + ∫ x in s, f x ∂(measure.dirac a) = if a ∈ s then f a else 0 := +begin + rw [restrict_dirac], + split_ifs, + { exact integral_dirac _ _, }, + { exact integral_zero_measure _, }, +end + lemma mul_meas_ge_le_integral_of_nonneg [is_finite_measure μ] {f : α → ℝ} (hf_nonneg : 0 ≤ f) (hf_int : integrable f μ) (ε : ℝ) : ε * (μ {x | ε ≤ f x}).to_real ≤ ∫ x, f x ∂μ := diff --git a/src/measure_theory/integral/lebesgue.lean b/src/measure_theory/integral/lebesgue.lean index 685b546bb933e..7f5ee430c2000 100644 --- a/src/measure_theory/integral/lebesgue.lean +++ b/src/measure_theory/integral/lebesgue.lean @@ -32,6 +32,46 @@ open_locale classical topology big_operators nnreal ennreal measure_theory namespace measure_theory + +section move_this + +variables {α : Type*} {mα : measurable_space α} {a : α} {s : set α} +include mα + +-- todo after the port: move to measure_theory/measure/measure_space +lemma restrict_dirac' (hs : measurable_set s) [decidable (a ∈ s)] : + (measure.dirac a).restrict s = if a ∈ s then measure.dirac a else 0 := +begin + ext1 t ht, + classical, + simp only [measure.restrict_apply ht, measure.dirac_apply' _ (ht.inter hs), set.indicator_apply, + set.mem_inter_iff, pi.one_apply], + by_cases has : a ∈ s, + { simp only [has, and_true, if_true], + split_ifs with hat, + { rw measure.dirac_apply_of_mem hat, }, + { simp only [measure.dirac_apply' _ ht, set.indicator_apply, hat, if_false], }, }, + { simp only [has, and_false, if_false, measure.coe_zero, pi.zero_apply], }, +end + +-- todo after the port: move to measure_theory/measure/measure_space +lemma restrict_dirac [measurable_singleton_class α] [decidable (a ∈ s)] : + (measure.dirac a).restrict s = if a ∈ s then measure.dirac a else 0 := +begin + ext1 t ht, + classical, + simp only [measure.restrict_apply ht, measure.dirac_apply _, set.indicator_apply, + set.mem_inter_iff, pi.one_apply], + by_cases has : a ∈ s, + { simp only [has, and_true, if_true], + split_ifs with hat, + { rw measure.dirac_apply_of_mem hat, }, + { simp only [measure.dirac_apply' _ ht, set.indicator_apply, hat, if_false], }, }, + { simp only [has, and_false, if_false, measure.coe_zero, pi.zero_apply], }, +end + +end move_this + local infixr ` →ₛ `:25 := simple_func variables {α β γ δ : Type*} @@ -1288,6 +1328,27 @@ lemma lintegral_dirac [measurable_singleton_class α] (a : α) (f : α → ℝ ∫⁻ a, f a ∂(dirac a) = f a := by simp [lintegral_congr_ae (ae_eq_dirac f)] +lemma set_lintegral_dirac' {a : α} {f : α → ℝ≥0∞} (hf : measurable f) + {s : set α} (hs : measurable_set s) [decidable (a ∈ s)] : + ∫⁻ x in s, f x ∂(measure.dirac a) = if a ∈ s then f a else 0 := +begin + rw [restrict_dirac' hs], + swap, { apply_instance, }, + split_ifs, + { exact lintegral_dirac' _ hf, }, + { exact lintegral_zero_measure _, }, +end + +lemma set_lintegral_dirac {a : α} (f : α → ℝ≥0∞) + (s : set α) [measurable_singleton_class α] [decidable (a ∈ s)] : + ∫⁻ x in s, f x ∂(measure.dirac a) = if a ∈ s then f a else 0 := +begin + rw [restrict_dirac], + split_ifs, + { exact lintegral_dirac _ _, }, + { exact lintegral_zero_measure _, }, +end + lemma lintegral_count' {f : α → ℝ≥0∞} (hf : measurable f) : ∫⁻ a, f a ∂count = ∑' a, f a := begin diff --git a/src/probability/kernel/basic.lean b/src/probability/kernel/basic.lean index 9a85dfb185e3f..e2068f99f01cc 100644 --- a/src/probability/kernel/basic.lean +++ b/src/probability/kernel/basic.lean @@ -347,6 +347,56 @@ instance is_markov_kernel_deterministic {f : α → β} (hf : measurable f) : is_markov_kernel (deterministic f hf) := ⟨λ a, by { rw deterministic_apply hf, apply_instance, }⟩ +lemma lintegral_deterministic' {f : β → ℝ≥0∞} {g : α → β} {a : α} + (hg : measurable g) (hf : measurable f) : + ∫⁻ x, f x ∂(kernel.deterministic g hg a) = f (g a) := +by rw [kernel.deterministic_apply, lintegral_dirac' _ hf] + +@[simp] +lemma lintegral_deterministic {f : β → ℝ≥0∞} {g : α → β} {a : α} + (hg : measurable g) [measurable_singleton_class β] : + ∫⁻ x, f x ∂(kernel.deterministic g hg a) = f (g a) := +by rw [kernel.deterministic_apply, lintegral_dirac (g a) f] + +lemma set_lintegral_deterministic' {f : β → ℝ≥0∞} {g : α → β} {a : α} + (hg : measurable g) (hf : measurable f) {s : set β} (hs : measurable_set s) + [decidable (g a ∈ s)] : + ∫⁻ x in s, f x ∂(kernel.deterministic g hg a) = if g a ∈ s then f (g a) else 0 := +by rw [kernel.deterministic_apply, set_lintegral_dirac' hf hs] + +@[simp] +lemma set_lintegral_deterministic {f : β → ℝ≥0∞} {g : α → β} {a : α} + (hg : measurable g) [measurable_singleton_class β] (s : set β) [decidable (g a ∈ s)] : + ∫⁻ x in s, f x ∂(kernel.deterministic g hg a) = if g a ∈ s then f (g a) else 0 := +by rw [kernel.deterministic_apply, set_lintegral_dirac f s] + +lemma integral_deterministic' {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] + [complete_space E] {f : β → E} {g : α → β} {a : α} + (hg : measurable g) (hf : strongly_measurable f) : + ∫ x, f x ∂(kernel.deterministic g hg a) = f (g a) := +by rw [kernel.deterministic_apply, integral_dirac' _ _ hf] + +@[simp] +lemma integral_deterministic {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] + [complete_space E] {f : β → E} {g : α → β} {a : α} + (hg : measurable g) [measurable_singleton_class β] : + ∫ x, f x ∂(kernel.deterministic g hg a) = f (g a) := +by rw [kernel.deterministic_apply, integral_dirac _ (g a)] + +lemma set_integral_deterministic' {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] + [complete_space E] {f : β → E} {g : α → β} {a : α} + (hg : measurable g) (hf : strongly_measurable f) {s : set β} (hs : measurable_set s) + [decidable (g a ∈ s)] : + ∫ x in s, f x ∂(kernel.deterministic g hg a) = if g a ∈ s then f (g a) else 0 := +by rw [kernel.deterministic_apply, set_integral_dirac' hf _ hs] + +@[simp] +lemma set_integral_deterministic {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] + [complete_space E] {f : β → E} {g : α → β} {a : α} + (hg : measurable g) [measurable_singleton_class β] (s : set β) [decidable (g a ∈ s)] : + ∫ x in s, f x ∂(kernel.deterministic g hg a) = if g a ∈ s then f (g a) else 0 := +by rw [kernel.deterministic_apply, set_integral_dirac f _ s] + end deterministic section const @@ -373,6 +423,28 @@ instance is_markov_kernel_const {μβ : measure β} [hμβ : is_probability_meas is_markov_kernel (const α μβ) := ⟨λ a, hμβ⟩ +@[simp] +lemma lintegral_const {f : β → ℝ≥0∞} {μ : measure β} {a : α} : + ∫⁻ x, f x ∂(kernel.const α μ a) = ∫⁻ x, f x ∂μ := +by rw kernel.const_apply + +@[simp] +lemma set_lintegral_const {f : β → ℝ≥0∞} {μ : measure β} {a : α} {s : set β} : + ∫⁻ x in s, f x ∂(kernel.const α μ a) = ∫⁻ x in s, f x ∂μ := +by rw kernel.const_apply + +@[simp] +lemma integral_const {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] [complete_space E] + {f : β → E} {μ : measure β} {a : α} : + ∫ x, f x ∂(kernel.const α μ a) = ∫ x, f x ∂μ := +by rw kernel.const_apply + +@[simp] +lemma set_integral_const {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] [complete_space E] + {f : β → E} {μ : measure β} {a : α} {s : set β} : + ∫ x in s, f x ∂(kernel.const α μ a) = ∫ x in s, f x ∂μ := +by rw kernel.const_apply + end const omit mα @@ -408,10 +480,27 @@ lemma restrict_apply' (κ : kernel α β) (hs : measurable_set s) (a : α) (ht : kernel.restrict κ hs a t = (κ a) (t ∩ s) := by rw [restrict_apply κ hs a, measure.restrict_apply ht] +@[simp] +lemma restrict_univ : kernel.restrict κ measurable_set.univ = κ := +by { ext1 a, rw [kernel.restrict_apply, measure.restrict_univ], } + +@[simp] lemma lintegral_restrict (κ : kernel α β) (hs : measurable_set s) (a : α) (f : β → ℝ≥0∞) : ∫⁻ b, f b ∂(kernel.restrict κ hs a) = ∫⁻ b in s, f b ∂(κ a) := by rw restrict_apply +@[simp] +lemma set_lintegral_restrict (κ : kernel α β) (hs : measurable_set s) (a : α) (f : β → ℝ≥0∞) + (t : set β) : + ∫⁻ b in t, f b ∂(kernel.restrict κ hs a) = ∫⁻ b in (t ∩ s), f b ∂(κ a) := +by rw [restrict_apply, measure.restrict_restrict' hs] + +@[simp] +lemma set_integral_restrict {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] + [complete_space E] {f : β → E} {a : α} (hs : measurable_set s) (t : set β) : + ∫ x in t, f x ∂(kernel.restrict κ hs a) = ∫ x in (t ∩ s), f x ∂(κ a) := +by rw [restrict_apply, measure.restrict_restrict' hs] + instance is_finite_kernel.restrict (κ : kernel α β) [is_finite_kernel κ] (hs : measurable_set s) : is_finite_kernel (kernel.restrict κ hs) := begin diff --git a/src/probability/kernel/composition.lean b/src/probability/kernel/composition.lean index 7177e33219350..625092688fc58 100644 --- a/src/probability/kernel/composition.lean +++ b/src/probability/kernel/composition.lean @@ -93,7 +93,7 @@ def comp_prod_fun (κ : kernel α β) (η : kernel (α × β) γ) (a : α) (s : lemma comp_prod_fun_empty (κ : kernel α β) (η : kernel (α × β) γ) (a : α) : comp_prod_fun κ η a ∅ = 0 := by simp only [comp_prod_fun, set.mem_empty_iff_false, set.set_of_false, measure_empty, - lintegral_const, zero_mul] + measure_theory.lintegral_const, zero_mul] lemma comp_prod_fun_Union (κ : kernel α β) (η : kernel (α × β) γ) [is_s_finite_kernel η] (a : α) (f : ℕ → set (β × γ)) (hf_meas : ∀ i, measurable_set (f i)) (hf_disj : pairwise (disjoint on f)) : @@ -336,7 +336,7 @@ begin let Cη := is_finite_kernel.bound η, calc ∫⁻ b, η (a, b) set.univ ∂(κ a) ≤ ∫⁻ b, Cη ∂(κ a) : lintegral_mono (λ b, measure_le_bound η (a, b) set.univ) - ... = Cη * κ a set.univ : lintegral_const Cη + ... = Cη * κ a set.univ : measure_theory.lintegral_const Cη ... = κ a set.univ * Cη : mul_comm _ _, end diff --git a/src/probability/kernel/invariance.lean b/src/probability/kernel/invariance.lean index ddc6ec188d05b..75ccfd8c63442 100644 --- a/src/probability/kernel/invariance.lean +++ b/src/probability/kernel/invariance.lean @@ -46,7 +46,7 @@ noncomputable def map_measure (κ : kernel α β) (μ : measure α) : measure β := measure.of_measurable (λ s hs, ∫⁻ x, κ x s ∂μ) - (by simp only [measure_empty, lintegral_const, zero_mul]) + (by simp only [measure_empty, measure_theory.lintegral_const, zero_mul]) begin intros f hf₁ hf₂, simp_rw [measure_Union hf₂ hf₁, From c596622fccd6e0321979d94931c964054dea2d26 Mon Sep 17 00:00:00 2001 From: Thomas Browning Date: Mon, 8 May 2023 17:26:30 +0000 Subject: [PATCH 0975/1291] feat(field_theory/intermediate_field): Inhabited instance (#18956) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit If `S : intermediate_field K L`, then `S →ₐ[K] L` is nonempty. Co-authored-by: Thomas Browning --- src/field_theory/intermediate_field.lean | 2 ++ 1 file changed, 2 insertions(+) diff --git a/src/field_theory/intermediate_field.lean b/src/field_theory/intermediate_field.lean index d23233382994f..de0b97b760c5a 100644 --- a/src/field_theory/intermediate_field.lean +++ b/src/field_theory/intermediate_field.lean @@ -352,6 +352,8 @@ S.to_subalgebra.range_val @[simp] lemma field_range_val : S.val.field_range = S := set_like.ext' subtype.range_val +instance alg_hom.inhabited : inhabited (S →ₐ[K] L) := ⟨S.val⟩ + lemma aeval_coe {R : Type*} [comm_ring R] [algebra R K] [algebra R L] [is_scalar_tower R K L] (x : S) (P : R[X]) : aeval (x : L) P = aeval x P := begin From d4437c68c8d350fc9d4e95e1e174409db35e30d7 Mon Sep 17 00:00:00 2001 From: Thomas Browning Date: Mon, 8 May 2023 17:26:31 +0000 Subject: [PATCH 0976/1291] feat(field_theory/adjoin,normal): `field_range` lemmas (#18959) This PR adds a couple useful lemmas regarding `field_range`. Co-authored-by: Thomas Browning --- src/field_theory/adjoin.lean | 8 ++++++++ src/field_theory/normal.lean | 9 +++++++++ 2 files changed, 17 insertions(+) diff --git a/src/field_theory/adjoin.lean b/src/field_theory/adjoin.lean index 91338e32a54dc..8d20609ac3665 100644 --- a/src/field_theory/adjoin.lean +++ b/src/field_theory/adjoin.lean @@ -198,6 +198,14 @@ lemma _root_.alg_hom.map_field_range {K L : Type*} [field K] [field L] [algebra (f : E →ₐ[F] K) (g : K →ₐ[F] L) : f.field_range.map g = (g.comp f).field_range := set_like.ext' (set.range_comp g f).symm +lemma _root_.alg_hom.field_range_eq_top {K : Type*} [field K] [algebra F K] {f : E →ₐ[F] K} : + f.field_range = ⊤ ↔ function.surjective f := +set_like.ext'_iff.trans set.range_iff_surjective + +@[simp] lemma _root_.alg_equiv.field_range_eq_top {K : Type*} [field K] [algebra F K] + (f : E ≃ₐ[F] K) : (f : E →ₐ[F] K).field_range = ⊤ := +alg_hom.field_range_eq_top.mpr f.surjective + end lattice section adjoin_def diff --git a/src/field_theory/normal.lean b/src/field_theory/normal.lean index 107129a8dbb4f..6907430ef736d 100644 --- a/src/field_theory/normal.lean +++ b/src/field_theory/normal.lean @@ -278,6 +278,15 @@ lemma alg_hom.restrict_normal_comp [normal F E] : alg_hom.ext (λ _, (algebra_map E K₃).injective (by simp only [alg_hom.comp_apply, alg_hom.restrict_normal_commutes])) +lemma alg_hom.field_range_of_normal {E : intermediate_field F K} [normal F E] (f : E →ₐ[F] K) : + f.field_range = E := +begin + haveI : is_scalar_tower F E E := by apply_instance, + let g := f.restrict_normal' E, + rw [←show E.val.comp ↑g = f, from fun_like.ext_iff.mpr (f.restrict_normal_commutes E), + ←alg_hom.map_field_range, g.field_range_eq_top, ←E.val.field_range_eq_map, E.field_range_val], +end + /-- Restrict algebra isomorphism to a normal subfield -/ def alg_equiv.restrict_normal [h : normal F E] : E ≃ₐ[F] E := alg_hom.restrict_normal' χ.to_alg_hom E From c1686dff26eaecf4efd4edd141ebf78de309ae80 Mon Sep 17 00:00:00 2001 From: Mauricio Collares Date: Mon, 8 May 2023 20:42:17 +0000 Subject: [PATCH 0977/1291] feat(data/real/ennreal): interactions between infi/supr and to_nnreal/to_real (#18946) Add lemmas relating `ennreal.to_real` and `ennreal.to_nnreal` to the indexed infimum. Note that a slightly different set of lemmas and proofs were added in leanprover-community/mathlib4#3457. Please make sure to switch to these versions when forward-porting, and add `#align`s to the existing lemmas. This also adds `inf_dist_eq_infi`, which is from the same mathlib4 PR. Co-authored-by: Eric Wieser --- src/data/real/ennreal.lean | 52 +++++++++++++++++++ .../metric_space/hausdorff_distance.lean | 8 +++ 2 files changed, 60 insertions(+) diff --git a/src/data/real/ennreal.lean b/src/data/real/ennreal.lean index 5d778bbb1602e..dabfc510a17e6 100644 --- a/src/data/real/ennreal.lean +++ b/src/data/real/ennreal.lean @@ -657,6 +657,13 @@ section complete_lattice lemma coe_Sup {s : set ℝ≥0} : bdd_above s → (↑(Sup s) : ℝ≥0∞) = (⨆a∈s, ↑a) := with_top.coe_Sup lemma coe_Inf {s : set ℝ≥0} : s.nonempty → (↑(Inf s) : ℝ≥0∞) = (⨅a∈s, ↑a) := with_top.coe_Inf +lemma coe_supr {ι : Sort*} {f : ι → ℝ≥0} (hf : bdd_above (range f)) : + (↑(supr f) : ℝ≥0∞) = ⨆a, ↑(f a) := +with_top.coe_supr _ hf +@[norm_cast] +lemma coe_infi {ι : Sort*} [nonempty ι] (f : ι → ℝ≥0) : (↑(infi f) : ℝ≥0∞) = (⨅ a, ↑(f a)) := +with_top.coe_infi f + lemma coe_mem_upper_bounds {s : set ℝ≥0} : ↑r ∈ upper_bounds ((coe : ℝ≥0 → ℝ≥0∞) '' s) ↔ r ∈ upper_bounds s := by simp [upper_bounds, ball_image_iff, -mem_image, *] {contextual := tt} @@ -1845,6 +1852,51 @@ end real section infi variables {ι : Sort*} {f g : ι → ℝ≥0∞} +lemma to_nnreal_infi (hf : ∀ i, f i ≠ ∞) : (infi f).to_nnreal = ⨅ i, (f i).to_nnreal := +begin + casesI is_empty_or_nonempty ι, + { rw [infi_of_empty, top_to_nnreal, nnreal.infi_empty] }, + { lift f to ι → ℝ≥0 using hf, + simp_rw [← coe_infi, to_nnreal_coe] }, +end + +lemma to_nnreal_Inf (s : set ℝ≥0∞) (hs : ∀ r ∈ s, r ≠ ∞) : + (Inf s).to_nnreal = Inf (ennreal.to_nnreal '' s) := +begin + have hf : ∀ i, (coe : s → ℝ≥0∞) i ≠ ∞ := λ ⟨r, rs⟩, hs r rs, + simpa only [←Inf_range, ←Inf_image', subtype.range_coe_subtype] using to_nnreal_infi hf +end + +lemma to_nnreal_supr (hf : ∀ i, f i ≠ ∞) : (supr f).to_nnreal = ⨆ i, (f i).to_nnreal := +begin + lift f to ι → ℝ≥0 using hf, + simp_rw to_nnreal_coe, + by_cases h : bdd_above (range f), + { rw [← coe_supr h, to_nnreal_coe] }, + { rw [nnreal.supr_of_not_bdd_above h, (with_top.supr_coe_eq_top f).mpr h, top_to_nnreal] } +end + +lemma to_nnreal_Sup (s : set ℝ≥0∞) (hs : ∀ r ∈ s, r ≠ ∞) : + (Sup s).to_nnreal = Sup (ennreal.to_nnreal '' s) := +begin + have hf : ∀ i, (coe : s → ℝ≥0∞) i ≠ ∞ := λ⟨r, rs⟩, hs r rs, + simpa only [←Sup_range, ←Sup_image', subtype.range_coe_subtype] using to_nnreal_supr hf +end + +lemma to_real_infi (hf : ∀ i, f i ≠ ∞) : (infi f).to_real = ⨅ i, (f i).to_real := +by simp only [ennreal.to_real, to_nnreal_infi hf, nnreal.coe_infi] + +lemma to_real_Inf (s : set ℝ≥0∞) (hf : ∀ r ∈ s, r ≠ ∞) : + (Inf s).to_real = Inf (ennreal.to_real '' s) := +by simp only [ennreal.to_real, to_nnreal_Inf s hf, nnreal.coe_Inf, set.image_image] + +lemma to_real_supr (hf : ∀ i, f i ≠ ∞) : (supr f).to_real = ⨆ i, (f i).to_real := +by simp only [ennreal.to_real, to_nnreal_supr hf, nnreal.coe_supr] + +lemma to_real_Sup (s : set ℝ≥0∞) (hf : ∀ r ∈ s, r ≠ ∞) : + (Sup s).to_real = Sup (ennreal.to_real '' s) := +by simp only [ennreal.to_real, to_nnreal_Sup s hf, nnreal.coe_Sup, set.image_image] + lemma infi_add : infi f + a = ⨅i, f i + a := le_antisymm (le_infi $ assume i, add_le_add (infi_le _ _) $ le_rfl) diff --git a/src/topology/metric_space/hausdorff_distance.lean b/src/topology/metric_space/hausdorff_distance.lean index a0b022d137142..39a3c9bc4a3ee 100644 --- a/src/topology/metric_space/hausdorff_distance.lean +++ b/src/topology/metric_space/hausdorff_distance.lean @@ -425,6 +425,14 @@ open emetric /-- The minimal distance of a point to a set -/ def inf_dist (x : α) (s : set α) : ℝ := ennreal.to_real (inf_edist x s) +theorem inf_dist_eq_infi : inf_dist x s = ⨅ y : s, dist x y := +begin + rw [inf_dist, inf_edist, infi_subtype', ennreal.to_real_infi], + { simp only [dist_edist], + refl }, + { exact λ _, edist_ne_top _ _ } +end + /-- the minimal distance is always nonnegative -/ lemma inf_dist_nonneg : 0 ≤ inf_dist x s := by simp [inf_dist] From 949dc57e616a621462062668c9f39e4e17b64b69 Mon Sep 17 00:00:00 2001 From: Eric Rodriguez Date: Mon, 8 May 2023 23:09:05 +0000 Subject: [PATCH 0978/1291] feat(*): generalise+add algebraic instances (#18947) These are needed for #18857 (splitting field diamond), and I feel as though it's getting unwieldy so having these separately for review may be nice. Co-authored-by: Eric Rodriguez <37984851+ericrbg@users.noreply.github.com> --- src/algebra/monoid_algebra/basic.lean | 5 +-- src/data/polynomial/basic.lean | 10 ++--- src/field_theory/normal.lean | 33 ++++++++------- src/ring_theory/adjoin_root.lean | 48 ++++++++++++++++++---- src/ring_theory/algebraic_independent.lean | 7 +--- src/ring_theory/ideal/quotient.lean | 19 +++++++-- 6 files changed, 83 insertions(+), 39 deletions(-) diff --git a/src/algebra/monoid_algebra/basic.lean b/src/algebra/monoid_algebra/basic.lean index a23033c2a43ba..b0df6344a58e8 100644 --- a/src/algebra/monoid_algebra/basic.lean +++ b/src/algebra/monoid_algebra/basic.lean @@ -295,9 +295,8 @@ instance [monoid R] [semiring k] [distrib_mul_action R k] [has_faithful_smul R k has_faithful_smul R (monoid_algebra k G) := finsupp.has_faithful_smul -instance [monoid R] [monoid S] [semiring k] [distrib_mul_action R k] [distrib_mul_action S k] - [has_smul R S] [is_scalar_tower R S k] : - is_scalar_tower R S (monoid_algebra k G) := +instance [semiring k] [smul_zero_class R k] [smul_zero_class S k] [has_smul R S] + [is_scalar_tower R S k] : is_scalar_tower R S (monoid_algebra k G) := finsupp.is_scalar_tower G k instance [monoid R] [monoid S] [semiring k] [distrib_mul_action R k] [distrib_mul_action S k] diff --git a/src/data/polynomial/basic.lean b/src/data/polynomial/basic.lean index 5bd63b93712af..09b1dd24f7c12 100644 --- a/src/data/polynomial/basic.lean +++ b/src/data/polynomial/basic.lean @@ -191,7 +191,7 @@ instance {S} [monoid S] [distrib_mul_action S R] : distrib_mul_action S R[X] := function.injective.distrib_mul_action ⟨to_finsupp, to_finsupp_zero, to_finsupp_add⟩ to_finsupp_injective to_finsupp_smul -instance {S} [monoid S] [distrib_mul_action S R] [has_faithful_smul S R] : +instance {S} [smul_zero_class S R] [has_faithful_smul S R] : has_faithful_smul S R[X] := { eq_of_smul_eq_smul := λ s₁ s₂ h, eq_of_smul_eq_smul $ λ a : ℕ →₀ R, congr_arg to_finsupp (h ⟨a⟩) } @@ -199,12 +199,12 @@ instance {S} [semiring S] [module S R] : module S R[X] := function.injective.module _ ⟨to_finsupp, to_finsupp_zero, to_finsupp_add⟩ to_finsupp_injective to_finsupp_smul -instance {S₁ S₂} [monoid S₁] [monoid S₂] [distrib_mul_action S₁ R] [distrib_mul_action S₂ R] +instance {S₁ S₂} [smul_zero_class S₁ R] [smul_zero_class S₂ R] [smul_comm_class S₁ S₂ R] : smul_comm_class S₁ S₂ R[X] := ⟨by { rintros _ _ ⟨⟩, simp_rw [←of_finsupp_smul, smul_comm] }⟩ -instance {S₁ S₂} [has_smul S₁ S₂] [monoid S₁] [monoid S₂] [distrib_mul_action S₁ R] - [distrib_mul_action S₂ R] [is_scalar_tower S₁ S₂ R] : is_scalar_tower S₁ S₂ R[X] := +instance {S₁ S₂} [has_smul S₁ S₂] [smul_zero_class S₁ R] [smul_zero_class S₂ R] + [is_scalar_tower S₁ S₂ R] : is_scalar_tower S₁ S₂ R[X] := ⟨by { rintros _ _ ⟨⟩, simp_rw [←of_finsupp_smul, smul_assoc] }⟩ instance is_scalar_tower_right {α K : Type*} [semiring K] [distrib_smul α K] @@ -212,7 +212,7 @@ instance is_scalar_tower_right {α K : Type*} [semiring K] [distrib_smul α K] ⟨by rintros _ ⟨⟩ ⟨⟩; simp_rw [smul_eq_mul, ← of_finsupp_smul, ← of_finsupp_mul, ← of_finsupp_smul, smul_mul_assoc]⟩ -instance {S} [monoid S] [distrib_mul_action S R] [distrib_mul_action Sᵐᵒᵖ R] +instance {S} [smul_zero_class S R] [smul_zero_class Sᵐᵒᵖ R] [is_central_scalar S R] : is_central_scalar S R[X] := ⟨by { rintros _ ⟨⟩, simp_rw [←of_finsupp_smul, op_smul_eq_smul] }⟩ diff --git a/src/field_theory/normal.lean b/src/field_theory/normal.lean index 6907430ef736d..213fe17544515 100644 --- a/src/field_theory/normal.lean +++ b/src/field_theory/normal.lean @@ -124,15 +124,22 @@ end lemma alg_equiv.transfer_normal (f : E ≃ₐ[F] E') : normal F E ↔ normal F E' := ⟨λ h, by exactI normal.of_alg_equiv f, λ h, by exactI normal.of_alg_equiv f.symm⟩ -lemma normal.of_is_splitting_field (p : F[X]) [hFEp : is_splitting_field F E p] : - normal F E := +-- seems to be causing a diamond in the below proof +-- however, this may be a fluke and the proof below uses non-canonical `algebra` instances: +-- when I replaced all the instances inside the proof with the "canonical" instances we have, +-- I had the (unprovable) goal (of the form) `adjoin_root.mk f (C x) = adjoin_root.mk f X` +-- for some `x, f`. So maybe this is indeed the correct approach and rewriting this proof is +-- salient in the future, or at least taking a closer look at the algebra instances it uses. +local attribute [-instance] adjoin_root.has_smul + +lemma normal.of_is_splitting_field (p : F[X]) [hFEp : is_splitting_field F E p] : normal F E := begin - by_cases hp : p = 0, - { haveI : is_splitting_field F F p, { rw hp, exact ⟨splits_zero _, subsingleton.elim _ _⟩ }, - exactI (alg_equiv.transfer_normal ((is_splitting_field.alg_equiv F p).trans - (is_splitting_field.alg_equiv E p).symm)).mp (normal_self F) }, + unfreezingI { rcases eq_or_ne p 0 with rfl | hp }, + { haveI : is_splitting_field F F 0 := ⟨splits_zero _, subsingleton.elim _ _⟩, + exact (alg_equiv.transfer_normal ((is_splitting_field.alg_equiv F 0).trans + (is_splitting_field.alg_equiv E 0).symm)).mp (normal_self F) }, refine normal_iff.2 (λ x, _), - haveI hFE : finite_dimensional F E := is_splitting_field.finite_dimensional E p, + have hFE : finite_dimensional F E := is_splitting_field.finite_dimensional E p, have Hx : is_integral F x := is_integral_of_noetherian (is_noetherian.iff_fg.2 hFE) x, refine ⟨Hx, or.inr _⟩, rintros q q_irred ⟨r, hr⟩, @@ -140,13 +147,11 @@ begin haveI := fact.mk q_irred, let pbED := adjoin_root.power_basis q_irred.ne_zero, haveI : finite_dimensional E D := power_basis.finite_dimensional pbED, - have finrankED : finite_dimensional.finrank E D = q.nat_degree := power_basis.finrank pbED, - letI : algebra F D := ring_hom.to_algebra ((algebra_map E D).comp (algebra_map F E)), - haveI : is_scalar_tower F E D := of_algebra_map_eq (λ _, rfl), + have finrankED : finite_dimensional.finrank E D = q.nat_degree, + { rw [power_basis.finrank pbED, adjoin_root.power_basis_dim] }, haveI : finite_dimensional F D := finite_dimensional.trans F E D, - suffices : nonempty (D →ₐ[F] E), - { cases this with ϕ, - rw [←with_bot.coe_one, degree_eq_iff_nat_degree_eq q_irred.ne_zero, ←finrankED], + rsuffices ⟨ϕ⟩ : nonempty (D →ₐ[F] E), + { rw [←with_bot.coe_one, degree_eq_iff_nat_degree_eq q_irred.ne_zero, ←finrankED], have nat_lemma : ∀ a b c : ℕ, a * b = c → c ≤ a → 0 < c → b = 1, { intros a b c h1 h2 h3, nlinarith }, exact nat_lemma _ _ _ (finite_dimensional.finrank_mul_finrank F E D) @@ -180,7 +185,7 @@ begin rw [←intermediate_field.to_subalgebra_le_to_subalgebra, intermediate_field.top_to_subalgebra], apply ge_trans (intermediate_field.algebra_adjoin_le_adjoin C S), suffices : (algebra.adjoin C S).restrict_scalars F - = (algebra.adjoin E {adjoin_root.root q}).restrict_scalars F, + = (algebra.adjoin E {adjoin_root.root q}).restrict_scalars F, { rw [adjoin_root.adjoin_root_eq_top, subalgebra.restrict_scalars_top, ←@subalgebra.restrict_scalars_top F C] at this, exact top_le_iff.mpr (subalgebra.restrict_scalars_injective F this) }, diff --git a/src/ring_theory/adjoin_root.lean b/src/ring_theory/adjoin_root.lean index 2539cba43fd04..d2c0007771cef 100644 --- a/src/ring_theory/adjoin_root.lean +++ b/src/ring_theory/adjoin_root.lean @@ -89,18 +89,41 @@ quotient.induction_on' x ih /-- Embedding of the original ring `R` into `adjoin_root f`. -/ def of : R →+* adjoin_root f := (mk f).comp C -instance [comm_semiring S] [algebra S R] : algebra S (adjoin_root f) := -ideal.quotient.algebra S +instance [distrib_smul S R] [is_scalar_tower S R R] : has_smul S (adjoin_root f) := +submodule.quotient.has_smul' _ -instance [comm_semiring S] [comm_semiring K] [has_smul S K] [algebra S R] [algebra K R] - [is_scalar_tower S K R] : - is_scalar_tower S K (adjoin_root f) := +instance [distrib_smul S R] [is_scalar_tower S R R] : distrib_smul S (adjoin_root f) := +submodule.quotient.distrib_smul' _ + +@[simp] +lemma smul_mk [distrib_smul S R] [is_scalar_tower S R R] (a : S) (x : R[X]) : + a • mk f x = mk f (a • x) := rfl + +lemma smul_of [distrib_smul S R] [is_scalar_tower S R R] (a : S) (x : R) : + a • of f x = of f (a • x) := +by rw [of, ring_hom.comp_apply, ring_hom.comp_apply, smul_mk, smul_C] + +instance (R₁ R₂ : Type*) [has_smul R₁ R₂] [distrib_smul R₁ R] [distrib_smul R₂ R] + [is_scalar_tower R₁ R R] [is_scalar_tower R₂ R R] [is_scalar_tower R₁ R₂ R] (f : R[X]) : + is_scalar_tower R₁ R₂ (adjoin_root f) := submodule.quotient.is_scalar_tower _ _ -instance [comm_semiring S] [comm_semiring K] [algebra S R] [algebra K R] [smul_comm_class S K R] : - smul_comm_class S K (adjoin_root f) := +instance (R₁ R₂ : Type*) [distrib_smul R₁ R] [distrib_smul R₂ R] + [is_scalar_tower R₁ R R] [is_scalar_tower R₂ R R] [smul_comm_class R₁ R₂ R] (f : R[X]) : + smul_comm_class R₁ R₂ (adjoin_root f) := submodule.quotient.smul_comm_class _ _ +instance is_scalar_tower_right [distrib_smul S R] [is_scalar_tower S R R] : + is_scalar_tower S (adjoin_root f) (adjoin_root f) := +ideal.quotient.is_scalar_tower_right + +instance [monoid S] [distrib_mul_action S R] [is_scalar_tower S R R] (f : R[X]) : + distrib_mul_action S (adjoin_root f) := +submodule.quotient.distrib_mul_action' _ + +instance [comm_semiring S] [algebra S R] : algebra S (adjoin_root f) := +ideal.quotient.algebra S + @[simp] lemma algebra_map_eq : algebra_map R (adjoin_root f) = of f := rfl variables (S) @@ -286,7 +309,16 @@ instance span_maximal_of_irreducible [fact (irreducible f)] : (span {f}).is_maxi principal_ideal_ring.is_maximal_of_irreducible $ fact.out _ noncomputable instance field [fact (irreducible f)] : field (adjoin_root f) := -{ ..adjoin_root.comm_ring f, +{ rat_cast := λ a, of f (a : K), + rat_cast_mk := λ a b h1 h2, + begin + letI : group_with_zero (adjoin_root f) := ideal.quotient.group_with_zero _, + rw [rat.cast_mk', _root_.map_mul, _root_.map_int_cast, map_inv₀, map_nat_cast], + end, + qsmul := (•), + qsmul_eq_mul' := λ a x, adjoin_root.induction_on _ x (λ p, + by { rw [smul_mk, of, ring_hom.comp_apply, ← (mk f).map_mul, polynomial.rat_smul_eq_C_mul] }), + ..adjoin_root.comm_ring f, ..ideal.quotient.field (span {f} : ideal K[X]) } lemma coe_injective (h : degree f ≠ 0) : function.injective (coe : K → adjoin_root f) := diff --git a/src/ring_theory/algebraic_independent.lean b/src/ring_theory/algebraic_independent.lean index 2cafbefcdae84..e7b6cc5ce00c3 100644 --- a/src/ring_theory/algebraic_independent.lean +++ b/src/ring_theory/algebraic_independent.lean @@ -411,12 +411,9 @@ lemma algebraic_independent.mv_polynomial_option_equiv_polynomial_adjoin_C hx.mv_polynomial_option_equiv_polynomial_adjoin (C r) = polynomial.C (algebra_map _ _ r) := begin - -- TODO: this instance is slow to infer - have h : is_scalar_tower R (mv_polynomial ι R) (polynomial (mv_polynomial ι R)) := - @polynomial.is_scalar_tower (mv_polynomial ι R) _ R _ _ _ _ _ _ _, rw [algebraic_independent.mv_polynomial_option_equiv_polynomial_adjoin_apply, aeval_C, - @is_scalar_tower.algebra_map_apply _ _ _ _ _ _ _ _ _ h, ← polynomial.C_eq_algebra_map, - polynomial.map_C, ring_hom.coe_coe, alg_equiv.commutes] + is_scalar_tower.algebra_map_apply R (mv_polynomial ι R), ← polynomial.C_eq_algebra_map, + polynomial.map_C, ring_hom.coe_coe, alg_equiv.commutes] end @[simp] diff --git a/src/ring_theory/ideal/quotient.lean b/src/ring_theory/ideal/quotient.lean index 4cefd306702b4..2b1256b650dcc 100644 --- a/src/ring_theory/ideal/quotient.lean +++ b/src/ring_theory/ideal/quotient.lean @@ -176,19 +176,30 @@ end open_locale classical -/-- quotient by maximal ideal is a field. def rather than instance, since users will have -computable inverses in some applications. +/-- The quotient by a maximal ideal is a group with zero. This is a `def` rather than `instance`, +since users will have computable inverses in some applications. + See note [reducible non-instances]. -/ @[reducible] -protected noncomputable def field (I : ideal R) [hI : I.is_maximal] : field (R ⧸ I) := +protected noncomputable def group_with_zero (I : ideal R) [hI : I.is_maximal] : + group_with_zero (R ⧸ I) := { inv := λ a, if ha : a = 0 then 0 else classical.some (exists_inv ha), mul_inv_cancel := λ a (ha : a ≠ 0), show a * dite _ _ _ = _, by rw dif_neg ha; exact classical.some_spec (exists_inv ha), inv_zero := dif_pos rfl, - ..quotient.comm_ring I, + ..(by apply_instance : monoid_with_zero (R ⧸ I)), ..quotient.is_domain I } +/-- The quotient by a maximal ideal is a field. This is a `def` rather than `instance`, since users +will have computable inverses (and `qsmul`, `rat_cast`) in some applications. + +See note [reducible non-instances]. -/ +@[reducible] +protected noncomputable def field (I : ideal R) [hI : I.is_maximal] : field (R ⧸ I) := +{ ..quotient.comm_ring I, + ..quotient.group_with_zero I } + /-- If the quotient by an ideal is a field, then the ideal is maximal. -/ theorem maximal_of_is_field (I : ideal R) (hqf : is_field (R ⧸ I)) : I.is_maximal := From 7d34004e19699895c13c86b78ae62bbaea0bc893 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Tue, 9 May 2023 05:07:50 +0000 Subject: [PATCH 0979/1291] chore(*): add mathlib4 synchronization comments (#18966) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.continued_fractions.computation.translations` * `analysis.normed.group.quotient` * `analysis.specific_limits.normed` * `computability.partrec` * `data.complex.orientation` * `data.sym.card` * `linear_algebra.orientation` * `measure_theory.tactic` * `topology.algebra.nonarchimedean.adic_topology` * `topology.category.Top.open_nhds` * `topology.continuous_function.algebra` --- src/algebra/continued_fractions/computation/translations.lean | 3 +++ src/analysis/normed/group/quotient.lean | 3 +++ src/analysis/specific_limits/normed.lean | 3 +++ src/computability/partrec.lean | 3 +++ src/data/complex/orientation.lean | 3 +++ src/data/sym/card.lean | 3 +++ src/linear_algebra/orientation.lean | 3 +++ src/measure_theory/tactic.lean | 3 +++ src/topology/algebra/nonarchimedean/adic_topology.lean | 3 +++ src/topology/category/Top/open_nhds.lean | 3 +++ src/topology/continuous_function/algebra.lean | 3 +++ 11 files changed, 33 insertions(+) diff --git a/src/algebra/continued_fractions/computation/translations.lean b/src/algebra/continued_fractions/computation/translations.lean index 2d7683d10376a..424eca916db3e 100644 --- a/src/algebra/continued_fractions/computation/translations.lean +++ b/src/algebra/continued_fractions/computation/translations.lean @@ -8,6 +8,9 @@ import algebra.continued_fractions.translations /-! # Basic Translation Lemmas Between Structures Defined for Computing Continued Fractions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Summary This is a collection of simple lemmas between the different structures used for the computation diff --git a/src/analysis/normed/group/quotient.lean b/src/analysis/normed/group/quotient.lean index 1c71d8f281869..310046e0dd45d 100644 --- a/src/analysis/normed/group/quotient.lean +++ b/src/analysis/normed/group/quotient.lean @@ -10,6 +10,9 @@ import ring_theory.ideal.quotient_operations /-! # Quotients of seminormed groups +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + For any `seminormed_add_comm_group M` and any `S : add_subgroup M`, we provide a `seminormed_add_comm_group`, the group quotient `M ⧸ S`. If `S` is closed, we provide `normed_add_comm_group (M ⧸ S)` (regardless of whether `M` itself is diff --git a/src/analysis/specific_limits/normed.lean b/src/analysis/specific_limits/normed.lean index 830407e577d14..291940cf21cc3 100644 --- a/src/analysis/specific_limits/normed.lean +++ b/src/analysis/specific_limits/normed.lean @@ -10,6 +10,9 @@ import analysis.specific_limits.basic /-! # A collection of specific limit computations +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains important specific limit computations in (semi-)normed groups/rings/spaces, as as well as such computations in `ℝ` when the natural proof passes through a fact about normed spaces. diff --git a/src/computability/partrec.lean b/src/computability/partrec.lean index f8124d3ee4546..67532caaa84e7 100644 --- a/src/computability/partrec.lean +++ b/src/computability/partrec.lean @@ -10,6 +10,9 @@ import data.pfun /-! # The partial recursive functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The partial recursive functions are defined similarly to the primitive recursive functions, but now all functions are partial, implemented using the `part` monad, and there is an additional operation, called diff --git a/src/data/complex/orientation.lean b/src/data/complex/orientation.lean index f2ba3d19f2737..074bfb08338d6 100644 --- a/src/data/complex/orientation.lean +++ b/src/data/complex/orientation.lean @@ -9,6 +9,9 @@ import linear_algebra.orientation /-! # The standard orientation on `ℂ`. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This had previously been in `linear_algebra.orientation`, but keeping it separate results in a significant import reduction. -/ diff --git a/src/data/sym/card.lean b/src/data/sym/card.lean index 213c8d23cb071..f45e2c505f828 100644 --- a/src/data/sym/card.lean +++ b/src/data/sym/card.lean @@ -10,6 +10,9 @@ import data.fintype.sum /-! # Stars and bars +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we prove (in `sym.card_sym_eq_multichoose`) that the function `multichoose n k` defined in `data/nat/choose/basic` counts the number of multisets of cardinality `k` over an alphabet of cardinality `n`. In conjunction with `nat.multichoose_eq` proved in diff --git a/src/linear_algebra/orientation.lean b/src/linear_algebra/orientation.lean index e706b7ccac3be..a17ec1101eba4 100644 --- a/src/linear_algebra/orientation.lean +++ b/src/linear_algebra/orientation.lean @@ -9,6 +9,9 @@ import linear_algebra.determinant /-! # Orientations of modules +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines orientations of modules. ## Main definitions diff --git a/src/measure_theory/tactic.lean b/src/measure_theory/tactic.lean index 8cc6cb66f517f..cbbc77040535d 100644 --- a/src/measure_theory/tactic.lean +++ b/src/measure_theory/tactic.lean @@ -11,6 +11,9 @@ import tactic.with_local_reducibility /-! # Tactics for measure theory +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Currently we have one domain-specific tactic for measure theory: `measurability`. This tactic is to a large extent a copy of the `continuity` tactic by Reid Barton. diff --git a/src/topology/algebra/nonarchimedean/adic_topology.lean b/src/topology/algebra/nonarchimedean/adic_topology.lean index e72e630fc6767..3190120464aac 100644 --- a/src/topology/algebra/nonarchimedean/adic_topology.lean +++ b/src/topology/algebra/nonarchimedean/adic_topology.lean @@ -12,6 +12,9 @@ import topology.algebra.uniform_ring /-! # Adic topology +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given a commutative ring `R` and an ideal `I` in `R`, this file constructs the unique topology on `R` which is compatible with the ring structure and such that a set is a neighborhood of zero if and only if it contains a power of `I`. This topology is non-archimedean: every diff --git a/src/topology/category/Top/open_nhds.lean b/src/topology/category/Top/open_nhds.lean index fffa7777873b7..e18f068a37e7a 100644 --- a/src/topology/category/Top/open_nhds.lean +++ b/src/topology/category/Top/open_nhds.lean @@ -8,6 +8,9 @@ import topology.category.Top.opens /-! # The category of open neighborhoods of a point +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given an object `X` of the category `Top` of topological spaces and a point `x : X`, this file builds the type `open_nhds x` of open neighborhoods of `x` in `X` and endows it with the partial order given by inclusion and the corresponding category structure (as a full subcategory of the diff --git a/src/topology/continuous_function/algebra.lean b/src/topology/continuous_function/algebra.lean index b40b6cb18e45a..f60f007b21629 100644 --- a/src/topology/continuous_function/algebra.lean +++ b/src/topology/continuous_function/algebra.lean @@ -18,6 +18,9 @@ import topology.uniform_space.compact_convergence /-! # Algebraic structures over continuous functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define instances of algebraic structures over the type `continuous_map α β` (denoted `C(α, β)`) of **bundled** continuous maps from `α` to `β`. For example, `C(α, β)` is a group when `β` is a group, a ring when `β` is a ring, etc. From cc5dd6244981976cc9da7afc4eee5682b037a013 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 9 May 2023 17:20:11 +0000 Subject: [PATCH 0980/1291] chore(representation_theory/group_cohomology): use `fin.cast_succ` instead of coe (#18950) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This replaces the cast from `fin n` to `fin (n + 1)` (defined as `⟨↑i % (n + 1), _⟩`) with `fin.cast_succ` (defined as `⟨↑i, _⟩`). This is the preferred spelling, and using it lets us simplify some proofs. This also removes the `g` argument from `partial_prod_right_inv`, as it was not used, and the interesting statement is the one without it. --- src/algebra/big_operators/fin.lean | 15 ++++++--------- .../group_cohomology/basic.lean | 4 ++-- .../group_cohomology/resolution.lean | 10 +++++----- 3 files changed, 13 insertions(+), 16 deletions(-) diff --git a/src/algebra/big_operators/fin.lean b/src/algebra/big_operators/fin.lean index 9b07547122eb8..4182c72eeeb8e 100644 --- a/src/algebra/big_operators/fin.lean +++ b/src/algebra/big_operators/fin.lean @@ -194,15 +194,14 @@ begin end) @[to_additive] lemma partial_prod_right_inv {G : Type*} [group G] - (g : G) (f : fin n → G) (i : fin n) : - ((g • partial_prod f) i)⁻¹ * (g • partial_prod f) i.succ = f i := + (f : fin n → G) (i : fin n) : + (partial_prod f i.cast_succ)⁻¹ * partial_prod f i.succ = f i := begin cases i with i hn, induction i with i hi generalizing hn, - { simp [←fin.succ_mk, partial_prod_succ] }, + { simp [-fin.succ_mk, partial_prod_succ] }, { specialize hi (lt_trans (nat.lt_succ_self i) hn), - simp only [mul_inv_rev, fin.coe_eq_cast_succ, fin.succ_mk, fin.cast_succ_mk, - smul_eq_mul, pi.smul_apply] at hi ⊢, + simp only [fin.coe_eq_cast_succ, fin.succ_mk, fin.cast_succ_mk] at hi ⊢, rw [←fin.succ_mk _ _ (lt_trans (nat.lt_succ_self _) hn), ←fin.succ_mk], simp only [partial_prod_succ, mul_inv_rev, fin.cast_succ_mk], assoc_rw [hi, inv_mul_cancel_left] } @@ -223,15 +222,13 @@ lemma inv_partial_prod_mul_eq_contract_nth {G : Type*} [group G] (partial_prod g (j.succ.succ_above k.cast_succ))⁻¹ * partial_prod g (j.succ_above k).succ = j.contract_nth has_mul.mul g k := begin - have := partial_prod_right_inv (1 : G) g, - simp only [one_smul, coe_eq_cast_succ] at this, rcases lt_trichotomy (k : ℕ) j with (h|h|h), - { rwa [succ_above_below, succ_above_below, this, contract_nth_apply_of_lt], + { rwa [succ_above_below, succ_above_below, partial_prod_right_inv, contract_nth_apply_of_lt], { assumption }, { rw [cast_succ_lt_iff_succ_le, succ_le_succ_iff, le_iff_coe_le_coe], exact le_of_lt h }}, { rwa [succ_above_below, succ_above_above, partial_prod_succ, cast_succ_fin_succ, ←mul_assoc, - this, contract_nth_apply_of_eq], + partial_prod_right_inv, contract_nth_apply_of_eq], { simpa only [le_iff_coe_le_coe, ←h] }, { rw [cast_succ_lt_iff_succ_le, succ_le_succ_iff, le_iff_coe_le_coe], exact le_of_eq h }}, diff --git a/src/representation_theory/group_cohomology/basic.lean b/src/representation_theory/group_cohomology/basic.lean index e63ce48b3af5c..a972d9f61f849 100644 --- a/src/representation_theory/group_cohomology/basic.lean +++ b/src/representation_theory/group_cohomology/basic.lean @@ -126,8 +126,8 @@ begin congr' 1, { congr, ext, - have := fin.partial_prod_right_inv (1 : G) g (fin.cast_succ x), - simp only [mul_inv_rev, fin.coe_eq_cast_succ, one_smul, fin.cast_succ_fin_succ] at *, + have := fin.partial_prod_right_inv g (fin.cast_succ x), + simp only [mul_inv_rev, fin.cast_succ_fin_succ] at *, rw [mul_assoc, ←mul_assoc _ _ (g x.succ), this, inv_mul_cancel_left] }, { exact finset.sum_congr rfl (λ j hj, by rw [diagonal_hom_equiv_symm_partial_prod_succ, fin.coe_succ]) } diff --git a/src/representation_theory/group_cohomology/resolution.lean b/src/representation_theory/group_cohomology/resolution.lean index d6049837d2f81..1d63d4ee38387 100644 --- a/src/representation_theory/group_cohomology/resolution.lean +++ b/src/representation_theory/group_cohomology/resolution.lean @@ -89,7 +89,7 @@ def Action_diagonal_succ (G : Type u) [group G] : Π (n : ℕ), (equiv.pi_fin_succ_above_equiv (λ j, G) 0).symm.to_iso (λ g, rfl)) lemma Action_diagonal_succ_hom_apply {G : Type u} [group G] {n : ℕ} (f : fin (n + 1) → G) : - (Action_diagonal_succ G n).hom.hom f = (f 0, λ i, (f i)⁻¹ * f i.succ) := + (Action_diagonal_succ G n).hom.hom f = (f 0, λ i, (f i.cast_succ)⁻¹ * f i.succ) := begin induction n with n hn, { exact prod.ext rfl (funext $ λ x, fin.elim0 x) }, @@ -100,7 +100,7 @@ begin left_regular_tensor_iso_hom_hom, tensor_iso_hom, mk_iso_hom_hom, equiv.to_iso_hom, tensor_hom, equiv.pi_fin_succ_above_equiv_symm_apply, tensor_apply, types_id_apply, tensor_rho, monoid_hom.one_apply, End.one_def, hn (λ (j : fin (n + 1)), f j.succ), - fin.coe_eq_cast_succ, fin.insert_nth_zero'], + fin.insert_nth_zero'], refine fin.cases (fin.cons_zero _ _) (λ i, _) x, { simp only [fin.cons_succ, mul_left_inj, inv_inj, fin.cast_succ_fin_succ], }}} end @@ -145,7 +145,7 @@ variables {k G n} lemma diagonal_succ_hom_single (f : Gⁿ⁺¹) (a : k) : (diagonal_succ k G n).hom.hom (single f a) = - single (f 0) 1 ⊗ₜ single (λ i, (f i)⁻¹ * f i.succ) a := + single (f 0) 1 ⊗ₜ single (λ i, (f i.cast_succ)⁻¹ * f i.succ) a := begin dunfold diagonal_succ, simpa only [iso.trans_hom, iso.symm_hom, Action.comp_hom, Module.comp_def, linear_map.comp_apply, @@ -268,7 +268,7 @@ inverse map sends a function `f : Gⁿ → A` to the representation morphism sen to `A`. -/ lemma diagonal_hom_equiv_symm_apply (f : (fin n → G) → A) (x : fin (n + 1) → G) : ((diagonal_hom_equiv n A).symm f).hom (finsupp.single x 1) - = A.ρ (x 0) (f (λ (i : fin n), (x ↑i)⁻¹ * x i.succ)) := + = A.ρ (x 0) (f (λ (i : fin n), (x i.cast_succ)⁻¹ * x i.succ)) := begin unfold diagonal_hom_equiv, simp only [linear_equiv.trans_symm, linear_equiv.symm_symm, linear_equiv.trans_apply, @@ -290,7 +290,7 @@ lemma diagonal_hom_equiv_symm_partial_prod_succ = f (fin.contract_nth a (*) g) := begin simp only [diagonal_hom_equiv_symm_apply, function.comp_app, fin.succ_succ_above_zero, - fin.partial_prod_zero, map_one, fin.coe_eq_cast_succ, fin.succ_succ_above_succ, + fin.partial_prod_zero, map_one, fin.succ_succ_above_succ, linear_map.one_apply, fin.partial_prod_succ], congr, ext, From 0013240bce820e3096cebb7ccf6d17e3f35f77ca Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Wed, 10 May 2023 08:33:30 +0000 Subject: [PATCH 0981/1291] chore(algebra/order/absolute_value): remove unneeded imports (#18978) This is just a CI verification backport of https://github.com/leanprover-community/mathlib4/pull/3869. Feel free to either close or merge! Co-authored-by: Scott Morrison --- src/algebra/order/absolute_value.lean | 2 -- 1 file changed, 2 deletions(-) diff --git a/src/algebra/order/absolute_value.lean b/src/algebra/order/absolute_value.lean index fef733c6dbf2e..ab4188ad4df84 100644 --- a/src/algebra/order/absolute_value.lean +++ b/src/algebra/order/absolute_value.lean @@ -6,8 +6,6 @@ Authors: Mario Carneiro, Anne Baanen import algebra.group_with_zero.units.lemmas import algebra.order.field.defs import algebra.order.hom.basic -import algebra.order.ring.abs -import algebra.ring.commute import algebra.ring.regular /-! From 0b20dffb0d2141f40a76ad71330dfadb0af75746 Mon Sep 17 00:00:00 2001 From: Michael Stoll Date: Wed, 10 May 2023 11:33:41 +0000 Subject: [PATCH 0982/1291] feat(number_theory/pell): add def and properties of fundamental solutions (#18901) This is the next step in developing the theory of Pell's equation. We define what a *fundamental solution* is, show that it exists and is unique and that it is characterized by the property that it (is positive and) generates the group of solutions up to sign. See [this Zulip thread](https://leanprover.zulipchat.com/#narrow/stream/217875-Is-there-code-for-X.3F/topic/Proving.20Pell's.20equation.20is.20solvable/near/343270338). --- src/number_theory/pell.lean | 319 +++++++++++++++++++++++++++++++++++- 1 file changed, 313 insertions(+), 6 deletions(-) diff --git a/src/number_theory/pell.lean b/src/number_theory/pell.lean index 126d65f2a2b8e..b8eae0615add6 100644 --- a/src/number_theory/pell.lean +++ b/src/number_theory/pell.lean @@ -30,9 +30,12 @@ $x^2 - d y^2 = 1$ has a nontrivial (i.e., with $y \ne 0$) solution in integers. See `pell.exists_of_not_is_square` and `pell.exists_nontrivial_of_not_is_square`. -The next step (TODO) will be to define the *fundamental solution* as the solution -with smallest $x$ among all solutions satisfying $x > 1$ and $y > 0$ and to show -that every solution is a power of the fundamental solution up to a (common) sign. +We then define the *fundamental solution* to be the solution +with smallest $x$ among all solutions satisfying $x > 1$ and $y > 0$. +We show that every solution is a power (in the sense of the group structure mentioned above) +of the fundamental solution up to a (common) sign, +see `pell.fundamental.eq_zpow_or_neg_zpow`, and that a (positive) solution has this property +if and only if it is fundamental, see `pell.pos_generator_iff_fundamental`. ## References @@ -45,8 +48,7 @@ Pell's equation ## TODO -* Provide the structure theory of the solution set to Pell's equation - and furthermore also for `x ^ 2 - d * y ^ 2 = -1` and further generalizations. +* Extend to `x ^ 2 - d * y ^ 2 = -1` and further generalizations. * Connect solutions to the continued fraction expansion of `√d`. -/ @@ -183,6 +185,23 @@ begin exact lt_irrefl _ (((one_lt_sq_iff $ zero_le_one.trans ha.le).mpr ha).trans_eq prop), end +/-- If a solution has `x > 1`, then `d` is positive. -/ +lemma d_pos_of_one_lt_x {a : solution₁ d} (ha : 1 < a.x) : 0 < d := +begin + refine pos_of_mul_pos_left _ (sq_nonneg a.y), + rw [a.prop_y, sub_pos], + exact one_lt_pow ha two_ne_zero, +end + +/-- If a solution has `x > 1`, then `d` is not a square. -/ +lemma d_nonsquare_of_one_lt_x {a : solution₁ d} (ha : 1 < a.x) : ¬ is_square d := +begin + have hp := a.prop, + rintros ⟨b, rfl⟩, + simp_rw [← sq, ← mul_pow, sq_sub_sq, int.mul_eq_one_iff_eq_one_or_neg_one] at hp, + rcases hp with ⟨hp₁, hp₂⟩ | ⟨hp₁, hp₂⟩; linarith [ha, hp₁, hp₂], +end + /-- A solution with `x = 1` is trivial. -/ lemma eq_one_of_x_eq_one (h₀ : d ≠ 0) {a : solution₁ d} (ha : a.x = 1) : a = 1 := begin @@ -245,6 +264,17 @@ begin exact y_mul_pos hax hay (x_pow_pos hax _) ih, } end +/-- If `(x, y)` is a solution with `x` and `y` positive, then all its powers with positive +exponents have positive `y`. -/ +lemma y_zpow_pos {a : solution₁ d} (hax : 0 < a.x) (hay : 0 < a.y) {n : ℤ} (hn : 0 < n) : + 0 < (a ^ n).y := +begin + lift n to ℕ using hn.le, + norm_cast at hn ⊢, + rw ← nat.succ_pred_eq_of_pos hn, + exact y_pow_succ_pos hax hay _, +end + /-- If `(x, y)` is a solution with `x` positive, then all its powers have positive `x`. -/ lemma x_zpow_pos {a : solution₁ d} (hax : 0 < a.x) (n : ℤ) : 0 < (a ^ n).x := begin @@ -258,7 +288,7 @@ end /-- If `(x, y)` is a solution with `x` and `y` positive, then the `y` component of any power has the same sign as the exponent. -/ lemma sign_y_zpow_eq_sign_of_x_pos_of_y_pos {a : solution₁ d} (hax : 0 < a.x) (hay : 0 < a.y) - (n : ℤ) : + (n : ℤ) : (a ^ n).y.sign = n.sign := begin rcases n with (_ | _) | _, @@ -398,4 +428,281 @@ end solution₁ end existence +/-! ### Fundamental solutions + +We define the notion of a *fundamental solution* of Pell's equation and +show that it exists and is unique (when `d` is positive and non-square) +and generates the group of solutions up to sign. +-/ + +variables {d : ℤ} + +/-- We define a solution to be *fundamental* if it has `x > 1` and `y > 0` +and its `x` is the smallest possible among solutions with `x > 1`. -/ +def is_fundamental (a : solution₁ d) : Prop := +1 < a.x ∧ 0 < a.y ∧ ∀ {b : solution₁ d}, 1 < b.x → a.x ≤ b.x + +namespace is_fundamental + +open solution₁ + +/-- A fundamental solution has positive `x`. -/ +lemma x_pos {a : solution₁ d} (h : is_fundamental a) : 0 < a.x := zero_lt_one.trans h.1 + +/-- If a fundamental solution exists, then `d` must be positive. -/ +lemma d_pos {a : solution₁ d} (h : is_fundamental a) : 0 < d := d_pos_of_one_lt_x h.1 + +/-- If a fundamental solution exists, then `d` must be a non-square. -/ +lemma d_nonsquare {a : solution₁ d} (h : is_fundamental a) : ¬ is_square d := +d_nonsquare_of_one_lt_x h.1 + +/-- If there is a fundamental solution, it is unique. -/ +lemma subsingleton {a b : solution₁ d} (ha : is_fundamental a) (hb : is_fundamental b) : a = b := +begin + have hx := le_antisymm (ha.2.2 hb.1) (hb.2.2 ha.1), + refine solution₁.ext hx _, + have : d * a.y ^ 2 = d * b.y ^ 2 := by rw [a.prop_y, b.prop_y, hx], + exact (sq_eq_sq ha.2.1.le hb.2.1.le).mp (int.eq_of_mul_eq_mul_left ha.d_pos.ne' this), +end + +/-- If `d` is positive and not a square, then a fundamental solution exists. -/ +lemma exists_of_not_is_square (h₀ : 0 < d) (hd : ¬ is_square d) : + ∃ a : solution₁ d, is_fundamental a := +begin + obtain ⟨a, ha₁, ha₂⟩ := exists_pos_of_not_is_square h₀ hd, + -- convert to `x : ℕ` to be able to use `nat.find` + have P : ∃ x' : ℕ, 1 < x' ∧ ∃ y' : ℤ, 0 < y' ∧ (x' : ℤ) ^ 2 - d * y' ^ 2 = 1, + { have hax := a.prop, + lift a.x to ℕ using (by positivity) with ax, + norm_cast at ha₁, + exact ⟨ax, ha₁, a.y, ha₂, hax⟩, }, + classical, -- to avoid having to show that the predicate is decidable + let x₁ := nat.find P, + obtain ⟨hx, y₁, hy₀, hy₁⟩ := nat.find_spec P, + refine ⟨mk x₁ y₁ hy₁, (by {rw x_mk, exact_mod_cast hx}), hy₀, λ b hb, _⟩, + rw x_mk, + have hb' := (int.to_nat_of_nonneg $ zero_le_one.trans hb.le).symm, + have hb'' := hb, + rw hb' at hb ⊢, + norm_cast at hb ⊢, + refine nat.find_min' P ⟨hb, |b.y|, abs_pos.mpr $ y_ne_zero_of_one_lt_x hb'', _⟩, + rw [← hb', sq_abs], + exact b.prop, +end + +/-- The map sending an integer `n` to the `y`-coordinate of `a^n` for a fundamental +solution `a` is stritcly increasing. -/ +lemma y_strict_mono {a : solution₁ d} (h : is_fundamental a) : + strict_mono (λ (n : ℤ), (a ^ n).y) := +begin + have H : ∀ (n : ℤ), 0 ≤ n → (a ^ n).y < (a ^ (n + 1)).y, + { intros n hn, + rw [← sub_pos, zpow_add, zpow_one, y_mul, add_sub_assoc], + rw show (a ^ n).y * a.x - (a ^ n).y = (a ^ n).y * (a.x - 1), from by ring, + refine add_pos_of_pos_of_nonneg (mul_pos (x_zpow_pos h.x_pos _) h.2.1) + (mul_nonneg _ (by {rw sub_nonneg, exact h.1.le})), + rcases hn.eq_or_lt with rfl | hn, + { simp only [zpow_zero, y_one], }, + { exact (y_zpow_pos h.x_pos h.2.1 hn).le, } }, + refine strict_mono_int_of_lt_succ (λ n, _), + cases le_or_lt 0 n with hn hn, + { exact H n hn, }, + { let m : ℤ := -n - 1, + have hm : n = -m - 1 := by simp only [neg_sub, sub_neg_eq_add, add_tsub_cancel_left], + rw [hm, sub_add_cancel, ← neg_add', zpow_neg, zpow_neg, y_inv, y_inv, neg_lt_neg_iff], + exact H _ (by linarith [hn]), } +end + +/-- If `a` is a fundamental solution, then `(a^m).y < (a^n).y` if and only if `m < n`. -/ +lemma zpow_y_lt_iff_lt {a : solution₁ d} (h : is_fundamental a) (m n : ℤ) : + (a ^ m).y < (a ^ n).y ↔ m < n := +begin + refine ⟨λ H, _, λ H, h.y_strict_mono H⟩, + contrapose! H, + exact h.y_strict_mono.monotone H, +end + +/-- The `n`th power of a fundamental solution is trivial if and only if `n = 0`. -/ +lemma zpow_eq_one_iff {a : solution₁ d} (h : is_fundamental a) (n : ℤ) : a ^ n = 1 ↔ n = 0 := +begin + rw ← zpow_zero a, + exact ⟨λ H, h.y_strict_mono.injective (congr_arg solution₁.y H), λ H, H ▸ rfl⟩, +end + +/-- A power of a fundamental solution is never equal to the negative of a power of this +fundamental solution. -/ +lemma zpow_ne_neg_zpow {a : solution₁ d} (h : is_fundamental a) {n n' : ℤ} : + a ^ n ≠ -a ^ n' := +begin + intro hf, + apply_fun solution₁.x at hf, + have H := x_zpow_pos h.x_pos n, + rw [hf, x_neg, lt_neg, neg_zero] at H, + exact lt_irrefl _ ((x_zpow_pos h.x_pos n').trans H), +end + +/-- The `x`-coordinate of a fundamental solution is a lower bound for the `x`-coordinate +of any positive solution. -/ +lemma x_le_x {a₁ : solution₁ d} (h : is_fundamental a₁) {a : solution₁ d} (hax : 1 < a.x) : + a₁.x ≤ a.x := +h.2.2 hax + +/-- The `y`-coordinate of a fundamental solution is a lower bound for the `y`-coordinate +of any positive solution. -/ +lemma y_le_y {a₁ : solution₁ d} (h : is_fundamental a₁) {a : solution₁ d} (hax : 1 < a.x) + (hay : 0 < a.y) : + a₁.y ≤ a.y := +begin + have H : d * (a₁.y ^ 2 - a.y ^ 2) = a₁.x ^ 2 - a.x ^ 2 := by { rw [a.prop_x, a₁.prop_x], ring }, + rw [← abs_of_pos hay, ← abs_of_pos h.2.1, ← sq_le_sq, ← mul_le_mul_left h.d_pos, ← sub_nonpos, + ← mul_sub, H, sub_nonpos, sq_le_sq, abs_of_pos (zero_lt_one.trans h.1), + abs_of_pos (zero_lt_one.trans hax)], + exact h.x_le_x hax, +end + +-- helper lemma for the next three results +lemma x_mul_y_le_y_mul_x {a₁ : solution₁ d} (h : is_fundamental a₁) {a : solution₁ d} + (hax : 1 < a.x) (hay : 0 < a.y) : + a.x * a₁.y ≤ a.y * a₁.x := +begin + rw [← abs_of_pos $ zero_lt_one.trans hax, ← abs_of_pos hay, ← abs_of_pos h.x_pos, + ← abs_of_pos h.2.1, ← abs_mul, ← abs_mul, ← sq_le_sq, mul_pow, mul_pow, a.prop_x, a₁.prop_x, + ← sub_nonneg], + ring_nf, + rw [sub_nonneg, sq_le_sq, abs_of_pos hay, abs_of_pos h.2.1], + exact h.y_le_y hax hay, +end + +/-- If we multiply a positive solution with the inverse of a fundamental solution, +the `y`-coordinate remains nonnegative. -/ +lemma mul_inv_y_nonneg {a₁ : solution₁ d} (h : is_fundamental a₁) {a : solution₁ d} (hax : 1 < a.x) + (hay : 0 < a.y) : + 0 ≤ (a * a₁⁻¹).y := +by simpa only [y_inv, mul_neg, y_mul, le_neg_add_iff_add_le, add_zero] + using h.x_mul_y_le_y_mul_x hax hay + +/-- If we multiply a positive solution with the inverse of a fundamental solution, +the `x`-coordinate stays positive. -/ +lemma mul_inv_x_pos {a₁ : solution₁ d} (h : is_fundamental a₁) {a : solution₁ d} (hax : 1 < a.x) + (hay : 0 < a.y) : + 0 < (a * a₁⁻¹).x := +begin + simp only [x_mul, x_inv, y_inv, mul_neg, lt_add_neg_iff_add_lt, zero_add], + refine (mul_lt_mul_left $ zero_lt_one.trans hax).mp _, + rw [(by ring : a.x * (d * (a.y * a₁.y)) = (d * a.y) * (a.x * a₁.y))], + refine ((mul_le_mul_left $ mul_pos h.d_pos hay).mpr $ x_mul_y_le_y_mul_x h hax hay).trans_lt _, + rw [← mul_assoc, mul_assoc d, ← sq, a.prop_y, ← sub_pos], + ring_nf, + exact zero_lt_one.trans h.1, +end + +/-- If we multiply a positive solution with the inverse of a fundamental solution, +the `x`-coordinate decreases. -/ +lemma mul_inv_x_lt_x {a₁ : solution₁ d} (h : is_fundamental a₁) {a : solution₁ d} (hax : 1 < a.x) + (hay : 0 < a.y) : + (a * a₁⁻¹).x < a.x := +begin + simp only [x_mul, x_inv, y_inv, mul_neg, add_neg_lt_iff_le_add'], + refine (mul_lt_mul_left h.2.1).mp _, + rw [(by ring : a₁.y * (a.x * a₁.x) = a.x * a₁.y * a₁.x)], + refine ((mul_le_mul_right $ zero_lt_one.trans h.1).mpr $ x_mul_y_le_y_mul_x h hax hay).trans_lt _, + rw [mul_assoc, ← sq, a₁.prop_x, ← sub_neg], + ring_nf, + rw [sub_neg, ← abs_of_pos hay, ← abs_of_pos h.2.1, ← abs_of_pos $ zero_lt_one.trans hax, + ← abs_mul, ← sq_lt_sq, mul_pow, a.prop_x], + calc + a.y ^ 2 = 1 * a.y ^ 2 : (one_mul _).symm + ... ≤ d * a.y ^ 2 : (mul_le_mul_right $ sq_pos_of_pos hay).mpr h.d_pos + ... < d * a.y ^ 2 + 1 : lt_add_one _ + ... = (1 + d * a.y ^ 2) * 1 : by rw [add_comm, mul_one] + ... ≤ (1 + d * a.y ^ 2) * a₁.y ^ 2 + : (mul_le_mul_left (by {have := h.d_pos, positivity})).mpr (sq_pos_of_pos h.2.1), +end + +/-- Any nonnegative solution is a power with nonnegative exponent of a fundamental solution. -/ +lemma eq_pow_of_nonneg {a₁ : solution₁ d} (h : is_fundamental a₁) {a : solution₁ d} (hax : 0 < a.x) + (hay : 0 ≤ a.y) : + ∃ n : ℕ, a = a₁ ^ n := +begin + lift a.x to ℕ using hax.le with ax hax', + induction ax using nat.strong_induction_on with x ih generalizing a, + cases hay.eq_or_lt with hy hy, + { -- case 1: `a = 1` + refine ⟨0, _⟩, + simp only [pow_zero], + ext; simp only [x_one, y_one], + { have prop := a.prop, + rw [← hy, sq (0 : ℤ), zero_mul, mul_zero, sub_zero, sq_eq_one_iff] at prop, + refine prop.resolve_right (λ hf, _), + have := (hax.trans_eq hax').le.trans_eq hf, + norm_num at this, }, + { exact hy.symm, } }, + { -- case 2: `a ≥ a₁` + have hx₁ : 1 < a.x := by nlinarith [a.prop, h.d_pos], + have hxx₁ := h.mul_inv_x_pos hx₁ hy, + have hxx₂ := h.mul_inv_x_lt_x hx₁ hy, + have hyy := h.mul_inv_y_nonneg hx₁ hy, + lift (a * a₁⁻¹).x to ℕ using hxx₁.le with x' hx', + obtain ⟨n, hn⟩ := ih x' (by exact_mod_cast hxx₂.trans_eq hax'.symm) hxx₁ hyy hx', + exact ⟨n + 1, by rw [pow_succ, ← hn, mul_comm a, ← mul_assoc, mul_inv_self, one_mul]⟩, }, +end + +/-- Every solution is, up to a sign, a power of a given fundamental solution. -/ +lemma eq_zpow_or_neg_zpow {a₁ : solution₁ d} (h : is_fundamental a₁) (a : solution₁ d) : + ∃ n : ℤ, a = a₁ ^ n ∨ a = -a₁ ^ n := +begin + obtain ⟨b, hbx, hby, hb⟩ := exists_pos_variant h.d_pos a, + obtain ⟨n, hn⟩ := h.eq_pow_of_nonneg hbx hby, + rcases hb with rfl | rfl | rfl | hb, + { exact ⟨n, or.inl (by exact_mod_cast hn)⟩, }, + { exact ⟨-n, or.inl (by simp [hn])⟩, }, + { exact ⟨n, or.inr (by simp [hn])⟩, }, + { rw set.mem_singleton_iff at hb, + rw hb, + exact ⟨-n, or.inr (by simp [hn])⟩, } +end + +end is_fundamental + +open solution₁ is_fundamental + +/-- When `d` is positive and not a square, then the group of solutions to the Pell equation +`x^2 - d*y^2 = 1` has a unique positive generator (up to sign). -/ +theorem exists_unique_pos_generator (h₀ : 0 < d) (hd : ¬ is_square d) : + ∃! a₁ : solution₁ d, 1 < a₁.x ∧ 0 < a₁.y ∧ ∀ a : solution₁ d, ∃ n : ℤ, a = a₁ ^ n ∨ a = -a₁ ^ n := +begin + obtain ⟨a₁, ha₁⟩ := is_fundamental.exists_of_not_is_square h₀ hd, + refine ⟨a₁, ⟨ha₁.1, ha₁.2.1, ha₁.eq_zpow_or_neg_zpow⟩, λ a (H : 1 < _ ∧ _), _⟩, + obtain ⟨Hx, Hy, H⟩ := H, + obtain ⟨n₁, hn₁⟩ := H a₁, + obtain ⟨n₂, hn₂⟩ := ha₁.eq_zpow_or_neg_zpow a, + rcases hn₂ with rfl | rfl, + { rw [← zpow_mul, eq_comm, @eq_comm _ a₁, ← mul_inv_eq_one, ← @mul_inv_eq_one _ _ _ a₁, + ← zpow_neg_one, neg_mul, ← zpow_add, ← sub_eq_add_neg] at hn₁, + cases hn₁, + { rcases int.is_unit_iff.mp (is_unit_of_mul_eq_one _ _ $ sub_eq_zero.mp $ + (ha₁.zpow_eq_one_iff (n₂ * n₁ - 1)).mp hn₁) with rfl | rfl, + { rw zpow_one, }, + { rw [zpow_neg_one, y_inv, lt_neg, neg_zero] at Hy, + exact false.elim (lt_irrefl _ $ ha₁.2.1.trans Hy), } }, + { rw [← zpow_zero a₁, eq_comm] at hn₁, + exact false.elim (ha₁.zpow_ne_neg_zpow hn₁), } }, + { rw [x_neg, lt_neg] at Hx, + have := (x_zpow_pos (zero_lt_one.trans ha₁.1) n₂).trans Hx, + norm_num at this, } +end + +/-- A positive solution is a generator (up to sign) of the group of all solutions to the +Pell equation `x^2 - d*y^2 = 1` if and only if it is a fundamental solution. -/ +theorem pos_generator_iff_fundamental (a : solution₁ d) : + (1 < a.x ∧ 0 < a.y ∧ ∀ b : solution₁ d, ∃ n : ℤ, b = a ^ n ∨ b = -a ^ n) ↔ is_fundamental a := +begin + refine ⟨λ h, _, λ H, ⟨H.1, H.2.1, H.eq_zpow_or_neg_zpow⟩⟩, + have h₀ := d_pos_of_one_lt_x h.1, + have hd := d_nonsquare_of_one_lt_x h.1, + obtain ⟨a₁, ha₁⟩ := is_fundamental.exists_of_not_is_square h₀ hd, + obtain ⟨b, hb₁, hb₂⟩ := exists_unique_pos_generator h₀ hd, + rwa [hb₂ a h, ← hb₂ a₁ ⟨ha₁.1, ha₁.2.1, ha₁.eq_zpow_or_neg_zpow⟩], +end + end pell From a07d750983b94c530ab69a726862c2ab6802b38c Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 10 May 2023 11:33:42 +0000 Subject: [PATCH 0983/1291] =?UTF-8?q?refactor(algebra):=20Redefine=20`a=20?= =?UTF-8?q?=E2=89=A1=20b=20[PMOD=20p]`=20(#18958)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit as `∃ n : ℤ, b - a = n • p` instead of `∀ n : ℤ, b - n • p ∉ set.Ioo a (a + p)`. Since this new definition doesn't require an order on `α`, we move it to a new file `algebra.modeq`. Expand the API. Co-authored-by: Eric Wieser --- src/algebra/group/basic.lean | 3 + src/algebra/group_power/lemmas.lean | 5 + src/algebra/hom/units.lean | 4 +- src/algebra/modeq.lean | 219 +++++++++++++++++++++++++ src/algebra/order/to_interval_mod.lean | 101 +++++------- src/topology/instances/add_circle.lean | 12 +- 6 files changed, 276 insertions(+), 68 deletions(-) create mode 100644 src/algebra/modeq.lean diff --git a/src/algebra/group/basic.lean b/src/algebra/group/basic.lean index bff1c7e7278ad..e4d064f28f2aa 100644 --- a/src/algebra/group/basic.lean +++ b/src/algebra/group/basic.lean @@ -451,6 +451,9 @@ by rw [div_eq_mul_inv, mul_right_inv a] lemma mul_div_cancel'' (a b : G) : a * b / b = a := by rw [div_eq_mul_inv, mul_inv_cancel_right a b] +@[simp, to_additive sub_add_cancel''] +lemma div_mul_cancel''' (a b : G) : a / (b * a) = b⁻¹ := by rw [←inv_div, mul_div_cancel''] + @[simp, to_additive] lemma mul_div_mul_right_eq_div (a b c : G) : (a * c) / (b * c) = a / b := by rw [div_mul_eq_div_div_swap]; simp only [mul_left_inj, eq_self_iff_true, mul_div_cancel''] diff --git a/src/algebra/group_power/lemmas.lean b/src/algebra/group_power/lemmas.lean index aad2d21b28729..c01093f2fd6bf 100644 --- a/src/algebra/group_power/lemmas.lean +++ b/src/algebra/group_power/lemmas.lean @@ -423,6 +423,11 @@ by { dsimp [bit1], rw [add_mul, bit0_mul, one_mul], } lemma mul_bit1 [non_assoc_ring R] {n r : R} : r * bit1 n = (2 : ℤ) • (r * n) + r := by { dsimp [bit1], rw [mul_add, mul_bit0, mul_one], } +/-- Note this holds in marginally more generality than `int.cast_mul` -/ +lemma int.cast_mul_eq_zsmul_cast [add_comm_group_with_one α] : ∀ m n, ((m * n : ℤ) : α) = m • n := +λ m, int.induction_on' m 0 (by simp) (λ k _ ih n, by simp [add_mul, add_zsmul, ih]) + (λ k _ ih n, by simp [sub_mul, sub_zsmul, ih, ←sub_eq_add_neg]) + @[simp] theorem zsmul_eq_mul [ring R] (a : R) : ∀ (n : ℤ), n • a = n * a | (n : ℕ) := by rw [coe_nat_zsmul, nsmul_eq_mul, int.cast_coe_nat] | -[1+ n] := by simp [nat.cast_succ, neg_add_rev, int.cast_neg_succ_of_nat, add_mul] diff --git a/src/algebra/hom/units.lean b/src/algebra/hom/units.lean index 00aff194746f0..c6e5cc4987288 100644 --- a/src/algebra/hom/units.lean +++ b/src/algebra/hom/units.lean @@ -278,7 +278,9 @@ h.eq_div_iff.2 @[to_additive] protected lemma div_eq_one_iff_eq (h : is_unit b) : a / b = 1 ↔ a = b := ⟨eq_of_div_eq_one, λ hab, hab.symm ▸ h.div_self⟩ -@[to_additive] protected lemma div_mul_left (h : is_unit b) : b / (a * b) = 1 / a := +/-- The `group` version of this lemma is `div_mul_cancel'''` -/ +@[to_additive "The `add_group` version of this lemma is `sub_add_cancel''`"] +protected lemma div_mul_left (h : is_unit b) : b / (a * b) = 1 / a := by rw [div_eq_mul_inv, mul_inv_rev, h.mul_inv_cancel_left, one_div] @[to_additive] protected lemma mul_div_mul_right (h : is_unit c) (a b : α) : diff --git a/src/algebra/modeq.lean b/src/algebra/modeq.lean new file mode 100644 index 0000000000000..26f397630e57b --- /dev/null +++ b/src/algebra/modeq.lean @@ -0,0 +1,219 @@ +/- +Copyright (c) 2023 Yaël Dillies. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies +-/ +import data.int.modeq +import group_theory.quotient_group + +/-! +# Equality modulo an element + +This file defines equality modulo an element in a commutative group. + +## Main definitions + +* `a ≡ b [PMOD p]`: `a` and `b` are congruent modulo a`p`. + +## See also + +`smodeq` is a generalisation to arbitrary submodules. + +## TODO + +Delete `int.modeq` in favour of `add_comm_group.modeq`. Generalise `smodeq` to `add_subgroup` and +redefine `add_comm_group.modeq` using it. Once this is done, we can rename `add_comm_group.modeq` +to `add_subgroup.modeq` and multiplicativise it. Longer term, we could generalise to submonoids and +also unify with `nat.modeq`. +-/ + +namespace add_comm_group +variables {α : Type*} + +section add_comm_group +variables [add_comm_group α] {p a a₁ a₂ b b₁ b₂ c : α} {n : ℕ} {z : ℤ} + +/-- `a ≡ b [PMOD p]` means that `b` is congruent to `a` modulo `p`. + +Equivalently (as shown in `algebra.order.to_interval_mod`), `b` does not lie in the open interval +`(a, a + p)` modulo `p`, or `to_Ico_mod hp a` disagrees with `to_Ioc_mod hp a` at `b`, or +`to_Ico_div hp a` disagrees with `to_Ioc_div hp a` at `b`. -/ +def modeq (p a b : α) : Prop := ∃ z : ℤ, b - a = z • p + +notation a ` ≡ `:50 b ` [PMOD `:50 p `]`:0 := modeq p a b + +@[refl, simp] lemma modeq_refl (a : α) : a ≡ a [PMOD p] := ⟨0, by simp⟩ + +lemma modeq_rfl : a ≡ a [PMOD p] := modeq_refl _ + +lemma modeq_comm : a ≡ b [PMOD p] ↔ b ≡ a [PMOD p] := +(equiv.neg _).exists_congr_left.trans $ by simp [modeq, ←neg_eq_iff_eq_neg] + +alias modeq_comm ↔ modeq.symm _ + +attribute [symm] modeq.symm + +@[trans] lemma modeq.trans : a ≡ b [PMOD p] → b ≡ c [PMOD p] → a ≡ c [PMOD p] := +λ ⟨m, hm⟩ ⟨n, hn⟩, ⟨m + n, by simp [add_smul, ←hm, ←hn]⟩ + +instance : is_refl _ (modeq p) := ⟨modeq_refl⟩ + +@[simp] lemma neg_modeq_neg : -a ≡ -b [PMOD p] ↔ a ≡ b [PMOD p] := +modeq_comm.trans $ by simp [modeq] + +alias neg_modeq_neg ↔ modeq.of_neg modeq.neg + +@[simp] lemma modeq_neg : a ≡ b [PMOD -p] ↔ a ≡ b [PMOD p] := +modeq_comm.trans $ by simp [modeq, ←neg_eq_iff_eq_neg] + +alias modeq_neg ↔ modeq.of_neg' modeq.neg' + +lemma modeq_sub (a b : α) : a ≡ b [PMOD b - a] := ⟨1, (one_smul _ _).symm⟩ + +@[simp] lemma modeq_zero : a ≡ b [PMOD 0] ↔ a = b := by simp [modeq, sub_eq_zero, eq_comm] + +@[simp] lemma self_modeq_zero : p ≡ 0 [PMOD p] := ⟨-1, by simp⟩ + +@[simp] lemma zsmul_modeq_zero (z : ℤ) : z • p ≡ 0 [PMOD p] := ⟨-z, by simp⟩ +lemma add_zsmul_modeq (z : ℤ) : a + z • p ≡ a [PMOD p] := ⟨-z, by simp⟩ +lemma zsmul_add_modeq (z : ℤ) : z • p + a ≡ a [PMOD p] := ⟨-z, by simp⟩ +lemma add_nsmul_modeq (n : ℕ) : a + n • p ≡ a [PMOD p] := ⟨-n, by simp⟩ +lemma nsmul_add_modeq (n : ℕ) : n • p + a ≡ a [PMOD p] := ⟨-n, by simp⟩ + +namespace modeq + +protected lemma add_zsmul (z : ℤ) : a ≡ b [PMOD p] → a + z • p ≡ b [PMOD p] := +(add_zsmul_modeq _).trans +protected lemma zsmul_add (z : ℤ) : a ≡ b [PMOD p] → z • p + a ≡ b [PMOD p] := +(zsmul_add_modeq _).trans +protected lemma add_nsmul (n : ℕ) : a ≡ b [PMOD p] → a + n • p ≡ b [PMOD p] := +(add_nsmul_modeq _).trans +protected lemma nsmul_add (n : ℕ) : a ≡ b [PMOD p] → n • p + a ≡ b [PMOD p] := +(nsmul_add_modeq _).trans + +protected lemma of_zsmul : a ≡ b [PMOD (z • p)] → a ≡ b [PMOD p] := +λ ⟨m, hm⟩, ⟨m * z, by rwa [mul_smul]⟩ + +protected lemma of_nsmul : a ≡ b [PMOD (n • p)] → a ≡ b [PMOD p] := +λ ⟨m, hm⟩, ⟨m * n, by rwa [mul_smul, coe_nat_zsmul]⟩ + +protected lemma zsmul : a ≡ b [PMOD p] → z • a ≡ z • b [PMOD (z • p)] := +Exists.imp $ λ m hm, by rw [←smul_sub, hm, smul_comm] + +protected lemma nsmul : a ≡ b [PMOD p] → n • a ≡ n • b [PMOD (n • p)] := +Exists.imp $ λ m hm, by rw [←smul_sub, hm, smul_comm] + +end modeq + +@[simp] lemma zsmul_modeq_zsmul [no_zero_smul_divisors ℤ α] (hn : z ≠ 0) : + z • a ≡ z • b [PMOD (z • p)] ↔ a ≡ b [PMOD p] := +exists_congr $ λ m, by rw [←smul_sub, smul_comm, smul_right_inj hn]; apply_instance + +@[simp] lemma nsmul_modeq_nsmul [no_zero_smul_divisors ℕ α] (hn : n ≠ 0) : + n • a ≡ n • b [PMOD (n • p)] ↔ a ≡ b [PMOD p] := +exists_congr $ λ m, by rw [←smul_sub, smul_comm, smul_right_inj hn]; apply_instance + +alias zsmul_modeq_zsmul ↔ modeq.zsmul_cancel _ +alias nsmul_modeq_nsmul ↔ modeq.nsmul_cancel _ + +namespace modeq + +@[simp] protected lemma add_iff_left : + a₁ ≡ b₁ [PMOD p] → (a₁ + a₂ ≡ b₁ + b₂ [PMOD p] ↔ a₂ ≡ b₂ [PMOD p]) := +λ ⟨m, hm⟩, (equiv.add_left m).symm.exists_congr_left.trans $ + by simpa [add_sub_add_comm, hm, add_smul] + +@[simp] protected lemma add_iff_right : + a₂ ≡ b₂ [PMOD p] → (a₁ + a₂ ≡ b₁ + b₂ [PMOD p] ↔ a₁ ≡ b₁ [PMOD p]) := +λ ⟨m, hm⟩, (equiv.add_right m).symm.exists_congr_left.trans $ + by simpa [add_sub_add_comm, hm, add_smul] + +@[simp] protected lemma sub_iff_left : + a₁ ≡ b₁ [PMOD p] → (a₁ - a₂ ≡ b₁ - b₂ [PMOD p] ↔ a₂ ≡ b₂ [PMOD p]) := +λ ⟨m, hm⟩, (equiv.sub_left m).symm.exists_congr_left.trans $ + by simpa [sub_sub_sub_comm, hm, sub_smul] + +@[simp] protected lemma sub_iff_right : + a₂ ≡ b₂ [PMOD p] → (a₁ - a₂ ≡ b₁ - b₂ [PMOD p] ↔ a₁ ≡ b₁ [PMOD p]) := +λ ⟨m, hm⟩, (equiv.sub_right m).symm.exists_congr_left.trans $ + by simpa [sub_sub_sub_comm, hm, sub_smul] + +alias modeq.add_iff_left ↔ add_left_cancel add +alias modeq.add_iff_right ↔ add_right_cancel _ +alias modeq.sub_iff_left ↔ sub_left_cancel sub +alias modeq.sub_iff_right ↔ sub_right_cancel _ + +attribute [protected] add_left_cancel add_right_cancel add sub_left_cancel sub_right_cancel sub + +protected lemma add_left (c : α) (h : a ≡ b [PMOD p]) : c + a ≡ c + b [PMOD p] := modeq_rfl.add h +protected lemma sub_left (c : α) (h : a ≡ b [PMOD p]) : c - a ≡ c - b [PMOD p] := modeq_rfl.sub h +protected lemma add_right (c : α) (h : a ≡ b [PMOD p]) : a + c ≡ b + c [PMOD p] := h.add modeq_rfl +protected lemma sub_right (c : α) (h : a ≡ b [PMOD p]) : a - c ≡ b - c [PMOD p] := h.sub modeq_rfl + +protected lemma add_left_cancel' (c : α) : c + a ≡ c + b [PMOD p] → a ≡ b [PMOD p] := +modeq_rfl.add_left_cancel + +protected lemma add_right_cancel' (c : α) : a + c ≡ b + c [PMOD p] → a ≡ b [PMOD p] := +modeq_rfl.add_right_cancel + +protected lemma sub_left_cancel' (c : α) : c - a ≡ c - b [PMOD p] → a ≡ b [PMOD p] := +modeq_rfl.sub_left_cancel + +protected lemma sub_right_cancel' (c : α) : a - c ≡ b - c [PMOD p] → a ≡ b [PMOD p] := +modeq_rfl.sub_right_cancel + +end modeq + +lemma modeq_sub_iff_add_modeq' : a ≡ b - c [PMOD p] ↔ c + a ≡ b [PMOD p] := by simp [modeq, sub_sub] +lemma modeq_sub_iff_add_modeq : a ≡ b - c [PMOD p] ↔ a + c ≡ b [PMOD p] := +modeq_sub_iff_add_modeq'.trans $ by rw add_comm +lemma sub_modeq_iff_modeq_add' : a - b ≡ c [PMOD p] ↔ a ≡ b + c [PMOD p] := +modeq_comm.trans $ modeq_sub_iff_add_modeq'.trans modeq_comm +lemma sub_modeq_iff_modeq_add : a - b ≡ c [PMOD p] ↔ a ≡ c + b [PMOD p] := +modeq_comm.trans $ modeq_sub_iff_add_modeq.trans modeq_comm + +@[simp] lemma sub_modeq_zero : a - b ≡ 0 [PMOD p] ↔ a ≡ b [PMOD p] := +by simp [sub_modeq_iff_modeq_add] + +@[simp] lemma add_modeq_left : a + b ≡ a [PMOD p] ↔ b ≡ 0 [PMOD p] := +by simp [←modeq_sub_iff_add_modeq'] + +@[simp] lemma add_modeq_right : a + b ≡ b [PMOD p] ↔ a ≡ 0 [PMOD p] := +by simp [←modeq_sub_iff_add_modeq] + +lemma modeq_iff_eq_add_zsmul : a ≡ b [PMOD p] ↔ ∃ z : ℤ, b = a + z • p := +by simp_rw [modeq, sub_eq_iff_eq_add'] + +lemma not_modeq_iff_ne_add_zsmul : ¬a ≡ b [PMOD p] ↔ ∀ z : ℤ, b ≠ a + z • p := +by rw [modeq_iff_eq_add_zsmul, not_exists] + +lemma modeq_iff_eq_mod_zmultiples : a ≡ b [PMOD p] ↔ (b : α ⧸ add_subgroup.zmultiples p) = a := +by simp_rw [modeq_iff_eq_add_zsmul, quotient_add_group.eq_iff_sub_mem, + add_subgroup.mem_zmultiples_iff, eq_sub_iff_add_eq', eq_comm] + +lemma not_modeq_iff_ne_mod_zmultiples : + ¬a ≡ b [PMOD p] ↔ (b : α ⧸ add_subgroup.zmultiples p) ≠ a := +modeq_iff_eq_mod_zmultiples.not + +end add_comm_group + +@[simp] lemma modeq_iff_int_modeq {a b z : ℤ} : a ≡ b [PMOD z] ↔ a ≡ b [ZMOD z] := +by simp [modeq, dvd_iff_exists_eq_mul_left, int.modeq_iff_dvd] + +section add_comm_group_with_one +variables [add_comm_group_with_one α] [char_zero α] + +@[simp, norm_cast] +lemma int_cast_modeq_int_cast {a b z : ℤ} : a ≡ b [PMOD (z : α)] ↔ a ≡ b [PMOD z] := +by simp_rw [modeq, ←int.cast_mul_eq_zsmul_cast]; norm_cast + +@[simp, norm_cast] +lemma nat_cast_modeq_nat_cast {a b n : ℕ} : a ≡ b [PMOD (n : α)] ↔ a ≡ b [MOD n] := +by simp_rw [←int.coe_nat_modeq_iff, ←modeq_iff_int_modeq, ←@int_cast_modeq_int_cast α, + int.cast_coe_nat] + +alias int_cast_modeq_int_cast ↔ modeq.of_int_cast modeq.int_cast +alias nat_cast_modeq_nat_cast ↔ _root_.nat.modeq.of_nat_cast modeq.nat_cast + +end add_comm_group_with_one +end add_comm_group diff --git a/src/algebra/order/to_interval_mod.lean b/src/algebra/order/to_interval_mod.lean index a8238c03d32f3..b547bcb2a117d 100644 --- a/src/algebra/order/to_interval_mod.lean +++ b/src/algebra/order/to_interval_mod.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 Joseph Myers. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Joseph Myers -/ +import algebra.modeq import algebra.module.basic import algebra.order.archimedean import algebra.periodic @@ -27,16 +28,12 @@ interval. * `to_Ioc_div hp a b` (where `hp : 0 < p`): The unique integer such that this multiple of `p`, subtracted from `b`, is in `Ioc a (a + p)`. * `to_Ioc_mod hp a b` (where `hp : 0 < p`): Reduce `b` to the interval `Ioc a (a + p)`. -* `a ≡ b [PMOD p]`: `a` and `b` are congruent modulo a multiple of `p`. See also `smodeq` which is a - more general version in arbitrary submodules. This is notation for `add_comm_group.modeq p a b`. - -## TODO - -Unify `smodeq` and `add_comm_group.modeq`, which were originally developed independently. -/ noncomputable theory +open add_comm_group + section linear_ordered_add_comm_group variables {α : Type*} [linear_ordered_add_comm_group α] [hα : archimedean α] {p : α} (hp : 0 < p) @@ -450,41 +447,49 @@ end /-! ### Links between the `Ico` and `Ioc` variants applied to the same element -/ section Ico_Ioc +variables {a b} namespace add_comm_group -variables (a b) -omit hα -/-- `add_comm_group.modeq p a b` means that `b` does not lie in the open interval `(a, a + p)` -modulo `p`. -Equivalently (as shown below), `b` is congruent to `a` modulo `p`, or `to_Ico_mod hp a` disagrees -with `to_Ioc_mod hp a` at `b`, or `to_Ico_div hp a` disagrees with `to_Ioc_div hp a` at `b`. -/ -def modeq (p a b : α) : Prop := ∀ z : ℤ, b - z • p ∉ set.Ioo a (a + p) -include hα +lemma modeq_iff_to_Ico_mod_eq_left : a ≡ b [PMOD p] ↔ to_Ico_mod hp a b = a := +modeq_iff_eq_add_zsmul.trans ⟨by { rintro ⟨n, rfl⟩, + rw [to_Ico_mod_add_zsmul, to_Ico_mod_apply_left] }, λ h, ⟨to_Ico_div hp a b, eq_add_of_sub_eq h⟩⟩ + +lemma modeq_iff_to_Ioc_mod_eq_right : a ≡ b [PMOD p] ↔ to_Ioc_mod hp a b = a + p := +begin + refine modeq_iff_eq_add_zsmul.trans ⟨_, λ h, ⟨to_Ioc_div hp a b + 1, _⟩⟩, + { rintro ⟨z, rfl⟩, + rw [to_Ioc_mod_add_zsmul, to_Ioc_mod_apply_left] }, + { rwa [add_one_zsmul, add_left_comm, ←sub_eq_iff_eq_add'] } +end -notation a ` ≡ `:50 b ` [PMOD `:50 p `]`:0 := modeq p a b +alias modeq_iff_to_Ico_mod_eq_left ↔ modeq.to_Ico_mod_eq_left _ +alias modeq_iff_to_Ioc_mod_eq_right ↔ modeq.to_Ico_mod_eq_right _ + +variables (a b) lemma tfae_modeq : - tfae [a ≡ b [PMOD p], + tfae [ + a ≡ b [PMOD p], + ∀ z : ℤ, b - z • p ∉ set.Ioo a (a + p), to_Ico_mod hp a b ≠ to_Ioc_mod hp a b, - to_Ico_mod hp a b + p = to_Ioc_mod hp a b, - to_Ico_mod hp a b = a] := + to_Ico_mod hp a b + p = to_Ioc_mod hp a b] := begin - rw modeq, - tfae_have : 2 → 1, + rw modeq_iff_to_Ico_mod_eq_left hp, + tfae_have : 3 → 2, { rw [←not_exists, not_imp_not], exact λ ⟨i, hi⟩, ((to_Ico_mod_eq_iff hp).2 ⟨set.Ioo_subset_Ico_self hi, i, (sub_add_cancel b _).symm⟩).trans ((to_Ioc_mod_eq_iff hp).2 ⟨set.Ioo_subset_Ioc_self hi, i, (sub_add_cancel b _).symm⟩).symm }, - tfae_have : 3 → 2, - { intro h, rw [←h, ne, eq_comm, add_right_eq_self], exact hp.ne' }, tfae_have : 4 → 3, + { intro h, rw [←h, ne, eq_comm, add_right_eq_self], exact hp.ne' }, + tfae_have : 1 → 4, { intro h, rw [h, eq_comm, to_Ioc_mod_eq_iff, set.right_mem_Ioc], refine ⟨lt_add_of_pos_right a hp, to_Ico_div hp a b - 1, _⟩, rw [sub_one_zsmul, add_add_add_comm, add_right_neg, add_zero], conv_lhs { rw [← to_Ico_mod_add_to_Ico_div_zsmul hp a b, h] } }, - tfae_have : 1 → 4, + tfae_have : 2 → 1, { rw [←not_exists, not_imp_comm], have h' := to_Ico_mod_mem_Ico hp a b, exact λ h, ⟨_, h'.1.lt_of_ne' h, h'.2⟩ }, @@ -493,25 +498,17 @@ end variables {a b} +lemma modeq_iff_not_forall_mem_Ioo_mod : + a ≡ b [PMOD p] ↔ ∀ z : ℤ, b - z • p ∉ set.Ioo a (a + p) := (tfae_modeq hp a b).out 0 1 lemma modeq_iff_to_Ico_mod_ne_to_Ioc_mod : - a ≡ b [PMOD p] ↔ to_Ico_mod hp a b ≠ to_Ioc_mod hp a b := (tfae_modeq hp a b).out 0 1 + a ≡ b [PMOD p] ↔ to_Ico_mod hp a b ≠ to_Ioc_mod hp a b := (tfae_modeq hp a b).out 0 2 lemma modeq_iff_to_Ico_mod_add_period_eq_to_Ioc_mod : - a ≡ b [PMOD p] ↔ to_Ico_mod hp a b + p = to_Ioc_mod hp a b := (tfae_modeq hp a b).out 0 2 -lemma modeq_iff_to_Ico_mod_eq_left : - a ≡ b [PMOD p] ↔ to_Ico_mod hp a b = a := (tfae_modeq hp a b).out 0 3 + a ≡ b [PMOD p] ↔ to_Ico_mod hp a b + p = to_Ioc_mod hp a b := (tfae_modeq hp a b).out 0 3 lemma not_modeq_iff_to_Ico_mod_eq_to_Ioc_mod : ¬a ≡ b [PMOD p] ↔ to_Ico_mod hp a b = to_Ioc_mod hp a b := (modeq_iff_to_Ico_mod_ne_to_Ioc_mod _).not_left -lemma modeq_iff_to_Ioc_mod_eq_right : a ≡ b [PMOD p] ↔ to_Ioc_mod hp a b = a + p := -begin - rw [modeq_iff_to_Ico_mod_ne_to_Ioc_mod hp, ne, to_Ico_mod_eq_iff hp, not_iff_comm], - obtain ⟨h₁, h₂⟩ := to_Ioc_mod_mem_Ioc hp a b, - exact ⟨λ h, ⟨⟨h₁.le, h₂.lt_of_ne h⟩, _, (to_Ioc_mod_add_to_Ioc_div_zsmul _ _ _).symm⟩, - λ h, h.1.2.ne⟩, -end - lemma not_modeq_iff_to_Ico_div_eq_to_Ioc_div : ¬a ≡ b [PMOD p] ↔ to_Ico_div hp a b = to_Ioc_div hp a b := by rw [not_modeq_iff_to_Ico_mod_eq_to_Ioc_mod hp, @@ -523,37 +520,21 @@ by rw [modeq_iff_to_Ico_mod_add_period_eq_to_Ioc_mod hp, to_Ico_mod, to_Ioc_mod, ← eq_sub_iff_add_eq, sub_sub, sub_right_inj, ← add_one_zsmul, (zsmul_strict_mono_left hp).injective.eq_iff] -include hp - -lemma modeq_iff_eq_add_zsmul : a ≡ b [PMOD p] ↔ ∃ z : ℤ, b = a + z • p := -begin - rw [modeq_iff_to_Ico_mod_eq_left hp], - split; intro h, - { rw ← h, - exact ⟨_, (to_Ico_mod_add_to_Ico_div_zsmul _ _ _).symm⟩ }, - { rw [to_Ico_mod_eq_iff, set.left_mem_Ico], - exact ⟨lt_add_of_pos_right a hp, h⟩, }, -end - -lemma not_modeq_iff_ne_add_zsmul : ¬a ≡ b [PMOD p] ↔ ∀ z : ℤ, b ≠ a + z • p := -by rw [modeq_iff_eq_add_zsmul hp, not_exists] +end add_comm_group -lemma modeq_iff_eq_mod_zmultiples : - a ≡ b [PMOD p] ↔ (b : α ⧸ add_subgroup.zmultiples p) = a := -by simp_rw [modeq_iff_eq_add_zsmul hp, quotient_add_group.eq_iff_sub_mem, - add_subgroup.mem_zmultiples_iff, eq_sub_iff_add_eq', eq_comm] +open add_comm_group -lemma not_modeq_iff_ne_mod_zmultiples : - ¬a ≡ b [PMOD p] ↔ (b : α ⧸ add_subgroup.zmultiples p) ≠ a := -(modeq_iff_eq_mod_zmultiples hp).not +/-- If `a` and `b` fall within the same cycle WRT `c`, then they are congruent modulo `p`. -/ +@[simp] lemma to_Ico_mod_inj {c : α} : to_Ico_mod hp c a = to_Ico_mod hp c b ↔ a ≡ b [PMOD p] := +by simp_rw [to_Ico_mod_eq_to_Ico_mod, modeq_iff_eq_add_zsmul, sub_eq_iff_eq_add'] -end add_comm_group +alias to_Ico_mod_inj ↔ _ add_comm_group.modeq.to_Ico_mod_eq_to_Ico_mod lemma Ico_eq_locus_Ioc_eq_Union_Ioo : {b | to_Ico_mod hp a b = to_Ioc_mod hp a b} = ⋃ z : ℤ, set.Ioo (a + z • p) (a + p + z • p) := begin ext1, simp_rw [set.mem_set_of, set.mem_Union, ← set.sub_mem_Ioo_iff_left, - ←add_comm_group.not_modeq_iff_to_Ico_mod_eq_to_Ioc_mod, add_comm_group.modeq, not_forall, + ←not_modeq_iff_to_Ico_mod_eq_to_Ioc_mod, modeq_iff_not_forall_mem_Ioo_mod hp, not_forall, not_not], end @@ -561,8 +542,8 @@ lemma to_Ioc_div_wcovby_to_Ico_div (a b : α) : to_Ioc_div hp a b ⩿ to_Ico_div begin suffices : to_Ioc_div hp a b = to_Ico_div hp a b ∨ to_Ioc_div hp a b + 1 = to_Ico_div hp a b, { rwa [wcovby_iff_eq_or_covby, ←order.succ_eq_iff_covby] }, - rw [eq_comm, ←add_comm_group.not_modeq_iff_to_Ico_div_eq_to_Ioc_div, - eq_comm, ←add_comm_group.modeq_iff_to_Ico_div_eq_to_Ioc_div_add_one], + rw [eq_comm, ←not_modeq_iff_to_Ico_div_eq_to_Ioc_div, + eq_comm, ←modeq_iff_to_Ico_div_eq_to_Ioc_div_add_one], exact em' _, end diff --git a/src/topology/instances/add_circle.lean b/src/topology/instances/add_circle.lean index 145643fd5f7a2..6847456e50aae 100644 --- a/src/topology/instances/add_circle.lean +++ b/src/topology/instances/add_circle.lean @@ -48,7 +48,7 @@ the rational circle `add_circle (1 : ℚ)`, and so we set things up more general noncomputable theory -open set function add_subgroup topological_space +open add_comm_group set function add_subgroup topological_space open_locale topology variables {𝕜 B : Type*} @@ -88,8 +88,7 @@ variables {x} (hx : (x : 𝕜 ⧸ zmultiples p) ≠ a) lemma to_Ico_mod_eventually_eq_to_Ioc_mod : to_Ico_mod hp a =ᶠ[𝓝 x] to_Ioc_mod hp a := is_open.mem_nhds (by {rw Ico_eq_locus_Ioc_eq_Union_Ioo, exact is_open_Union (λ i, is_open_Ioo)}) $ - (add_comm_group.not_modeq_iff_to_Ico_mod_eq_to_Ioc_mod hp).1 $ - (add_comm_group.not_modeq_iff_ne_mod_zmultiples hp).2 hx + (not_modeq_iff_to_Ico_mod_eq_to_Ioc_mod hp).1 $ not_modeq_iff_ne_mod_zmultiples.2 hx lemma continuous_at_to_Ico_mod : continuous_at (to_Ico_mod hp a) x := let h := to_Ico_mod_eventually_eq_to_Ioc_mod hp a hx in continuous_at_iff_continuous_left_right.2 $ @@ -498,11 +497,10 @@ lemma equiv_Icc_quot_comp_mk_eq_to_Ioc_mod : equiv_Icc_quot p a ∘ quotient.mk' λ x, quot.mk _ ⟨to_Ioc_mod hp.out a x, Ioc_subset_Icc_self $ to_Ioc_mod_mem_Ioc _ _ x⟩ := begin rw equiv_Icc_quot_comp_mk_eq_to_Ico_mod, funext, - by_cases add_comm_group.modeq p a x, - { simp_rw [(add_comm_group.modeq_iff_to_Ico_mod_eq_left hp.out).1 h, - (add_comm_group.modeq_iff_to_Ioc_mod_eq_right hp.out).1 h], + by_cases a ≡ x [PMOD p], + { simp_rw [(modeq_iff_to_Ico_mod_eq_left hp.out).1 h, (modeq_iff_to_Ioc_mod_eq_right hp.out).1 h], exact quot.sound endpoint_ident.mk }, - { simp_rw (add_comm_group.not_modeq_iff_to_Ico_mod_eq_to_Ioc_mod hp.out).1 h }, + { simp_rw (not_modeq_iff_to_Ico_mod_eq_to_Ioc_mod hp.out).1 h } end /-- The natural map from `[a, a + p] ⊂ 𝕜` with endpoints identified to `𝕜 / ℤ • p`, as a From e25a317463bd37d88e33da164465d8c47922b1cd Mon Sep 17 00:00:00 2001 From: Michael Stoll Date: Wed, 10 May 2023 14:51:29 +0000 Subject: [PATCH 0984/1291] feat(number_theory/diophantine_approximation): add Legendre's thm on rat'l approximations (#18460) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This adds *Legendre's Theorem* on rational approximations: ```lean lemma ex_continued_fraction_convergent_eq_rat {ξ : ℝ} {q : ℚ} (h : |ξ - q| < 1 / (2 * q.denom ^ 2)) : ∃ n, (generalized_continued_fraction.of ξ).convergents n = q ``` See this [Zulip thread](https://leanprover.zulipchat.com/#narrow/stream/116395-maths/topic/Diophantine.20approximation/near/323758115). Co-authored-by: Michael Stoll <99838730+MichaelStollBayreuth@users.noreply.github.com> --- .../diophantine_approximation.lean | 341 +++++++++++++++++- 1 file changed, 334 insertions(+), 7 deletions(-) diff --git a/src/number_theory/diophantine_approximation.lean b/src/number_theory/diophantine_approximation.lean index 0eb722539cf24..556dc53c8297a 100644 --- a/src/number_theory/diophantine_approximation.lean +++ b/src/number_theory/diophantine_approximation.lean @@ -3,18 +3,29 @@ Copyright (c) 2022 Michael Stoll. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Michael Geißer, Michael Stoll -/ -import tactic.basic +import algebra.continued_fractions.computation.approximation_corollaries +import algebra.continued_fractions.computation.translations +import combinatorics.pigeonhole +import data.int.units import data.real.irrational +import ring_theory.coprime.lemmas +import tactic.basic /-! # Diophantine Approximation -This file gives proofs of various versions of **Dirichlet's approximation theorem** -and its important consequence that when `ξ` is an irrational real number, then there are -infinitely many rationals `x/y` (in lowest terms) such that `|ξ - x/y| < 1/y^2`. - +The first part of this file gives proofs of various versions of +**Dirichlet's approximation theorem** and its important consequence that when $\xi$ is an +irrational real number, then there are infinitely many rationals $x/y$ (in lowest terms) +such that +$$\left|\xi - \frac{x}{y}\right| < \frac{1}{y^2} \,.$$ The proof is based on the pigeonhole principle. +The second part of the file gives a proof of **Legendre's Theorem** on rational approximation, +which states that if $\xi$ is a real number and $x/y$ is a rational number such that +$$\left|\xi - \frac{x}{y}\right| < \frac{1}{2y^2} \,,$$ +then $x/y$ must be a convergent of the continued fraction expansion of $\xi$. + ## Main statements The main results are three variants of Dirichlet's approximation theorem: @@ -36,18 +47,27 @@ We also show a converse, Both statements are combined to give an equivalence, `real.infinite_rat_abs_sub_lt_one_div_denom_sq_iff_irrational`. +There are two versions of Legendre's Theorem. One, `real.exists_rat_eq_convergent`, uses +`real.convergent`, a simple recursive definition of the convergents that is also defined +in this file, whereas the other, `real.exists_continued_fraction_convergent_eq_rat`, uses +`generalized_continued_fraction.convergents` of `generalized_continued_fraction.of ξ`. + ## Implementation notes We use the namespace `real` for the results on real numbers and `rat` for the results -on rational numbers. +on rational numbers. We introduce a secondary namespace `real.contfrac_legendre` +to separate off a definition and some technical auxiliary lemmas used in the proof +of Legendre's Theorem. For remarks on the proof of Legendre's Theorem, see below. ## References + (The German Wikipedia page on continued +fractions is much more extensive than the English one.) ## Tags -Diophantine approximation, Dirichlet's approximation theorem +Diophantine approximation, Dirichlet's approximation theorem, continued fraction -/ namespace real @@ -264,3 +284,310 @@ begin rw [H, (by push_cast : (1 : ℝ) / q.denom ^ 2 = (1 / q.denom ^ 2 : ℚ))], norm_cast, end + +/-! +### Legendre's Theorem on Rational Approximation + +We prove **Legendre's Theorem** on rational approximation: If $\xi$ is a real number and +$x/y$ is a rational number such that $|\xi - x/y| < 1/(2y^2)$, +then $x/y$ is a convergent of the continued fraction expansion of $\xi$. + +The proof is by induction. However, the induction proof does not work with the +statement as given, since the assumption is too weak to imply the corresponding +statement for the application of the induction hypothesis. This can be remedied +by making the statement slightly stronger. Namely, we assume that $|\xi - x/y| < 1/(y(2y-1))$ +when $y \ge 2$ and $-\frac{1}{2} < \xi - x < 1$ when $y = 1$. +-/ + +section convergent + +namespace real + +open int + +/-! +### Convergents: definition and API lemmas +-/ + +/-- We give a direct recursive definition of the convergents of the continued fraction +expansion of a real number `ξ`. The main reason for that is that we want to have the +convergents as rational numbers; the versions +`(generalized_continued_fraction.of ξ).convergents` and +`(generalized_continued_fraction.of ξ).convergents'` always give something of the +same type as `ξ`. We can then also use dot notation `ξ.convergent n`. +Another minor reason is that this demonstrates that the proof +of Legendre's theorem does not need anything beyond this definition. +We provide a proof that this definition agrees with the other one; +see `real.continued_fraction_convergent_eq_convergent`. +(Note that we use the fact that `1/0 = 0` here to make it work for rational `ξ`.) -/ +noncomputable def convergent : ℝ → ℕ → ℚ +| ξ 0 := ⌊ξ⌋ +| ξ (n + 1) := ⌊ξ⌋ + (convergent (fract ξ)⁻¹ n)⁻¹ + +/-- The zeroth convergent of `ξ` is `⌊ξ⌋`. -/ +@[simp] +lemma convergent_zero (ξ : ℝ) : ξ.convergent 0 = ⌊ξ⌋ := rfl + +/-- The `(n+1)`th convergent of `ξ` is the `n`th convergent of `1/(fract ξ)`. -/ +@[simp] +lemma convergent_succ (ξ : ℝ) (n : ℕ) : + ξ.convergent (n + 1) = ⌊ξ⌋ + ((fract ξ)⁻¹.convergent n)⁻¹ := +by simp only [convergent] + +/-- All convergents of `0` are zero. -/ +@[simp] +lemma convergent_of_zero (n : ℕ) : convergent 0 n = 0 := +begin + induction n with n ih, + { simp only [convergent_zero, floor_zero, cast_zero], }, + { simp only [ih, convergent_succ, floor_zero, cast_zero, fract_zero, add_zero, inv_zero], } +end + +/-- If `ξ` is an integer, all its convergents equal `ξ`. -/ +@[simp] +lemma convergent_of_int {ξ : ℤ} (n : ℕ) : convergent ξ n = ξ := +begin + cases n, + { simp only [convergent_zero, floor_int_cast], }, + { simp only [convergent_succ, floor_int_cast, fract_int_cast, convergent_of_zero, add_zero, + inv_zero], } +end + +/-! +Our `convergent`s agree with `generalized_continued_fraction.convergents`. +-/ + +open generalized_continued_fraction + +/-- The `n`th convergent of the `generalized_continued_fraction.of ξ` +agrees with `ξ.convergent n`. -/ +lemma continued_fraction_convergent_eq_convergent (ξ : ℝ) (n : ℕ) : + (generalized_continued_fraction.of ξ).convergents n = ξ.convergent n := +begin + induction n with n ih generalizing ξ, + { simp only [zeroth_convergent_eq_h, of_h_eq_floor, convergent_zero, rat.cast_coe_int], }, + { rw [convergents_succ, ih (fract ξ)⁻¹, convergent_succ, one_div], + norm_cast, } +end + +end real + +end convergent + +/-! +### The key technical condition for the induction proof +-/ + +namespace real + +open int + +/-- Define the technical condition to be used as assumption in the inductive proof. -/ +-- this is not `private`, as it is used in the public `exists_rat_eq_convergent'` below. +def contfrac_legendre.ass (ξ : ℝ) (u v : ℤ) : Prop := +is_coprime u v ∧ (v = 1 → (-(1 / 2) : ℝ) < ξ - u) ∧ |ξ - u / v| < (v * (2 * v - 1))⁻¹ + +-- ### Auxiliary lemmas + +-- This saves a few lines below, as it is frequently needed. +private lemma aux₀ {v : ℤ} (hv : 0 < v) : (0 : ℝ) < v ∧ (0 : ℝ) < 2 * v - 1 := +⟨cast_pos.mpr hv, by {norm_cast, linarith}⟩ + +-- In the following, we assume that `ass ξ u v` holds and `v ≥ 2`. + +variables {ξ : ℝ} {u v : ℤ} (hv : 2 ≤ v) (h : contfrac_legendre.ass ξ u v) +include hv h + +-- The fractional part of `ξ` is positive. +private lemma aux₁ : 0 < fract ξ := +begin + have hv₀ : (0 : ℝ) < v := cast_pos.mpr (zero_lt_two.trans_le hv), + obtain ⟨hv₁, hv₂⟩ := aux₀ (zero_lt_two.trans_le hv), + obtain ⟨hcop, _, h⟩ := h, + refine fract_pos.mpr (λ hf, _), + rw [hf] at h, + have H : (2 * v - 1 : ℝ) < 1, + { refine (mul_lt_iff_lt_one_right hv₀).mp + ((inv_lt_inv hv₀ (mul_pos hv₁ hv₂)).mp (lt_of_le_of_lt _ h)), + have h' : (⌊ξ⌋ : ℝ) - u / v = (⌊ξ⌋ * v - u) / v := by field_simp [hv₀.ne'], + rw [h', abs_div, abs_of_pos hv₀, ← one_div, div_le_div_right hv₀], + norm_cast, + rw [← zero_add (1 : ℤ), add_one_le_iff, abs_pos, sub_ne_zero], + rintro rfl, + cases is_unit_iff.mp (is_coprime_self.mp (is_coprime.mul_left_iff.mp hcop).2); linarith, }, + norm_cast at H, + linarith only [hv, H], +end + +-- An auxiliary lemma for the inductive step. +private lemma aux₂ : 0 < u - ⌊ξ⌋ * v ∧ u - ⌊ξ⌋ * v < v := +begin + obtain ⟨hcop, _, h⟩ := h, + obtain ⟨hv₀, hv₀'⟩ := aux₀ (zero_lt_two.trans_le hv), + have hv₁ : 0 < 2 * v - 1 := by linarith only [hv], + rw [← one_div, lt_div_iff (mul_pos hv₀ hv₀'), ← abs_of_pos (mul_pos hv₀ hv₀'), ← abs_mul, + sub_mul, ← mul_assoc, ← mul_assoc, div_mul_cancel _ hv₀.ne', abs_sub_comm, abs_lt, + lt_sub_iff_add_lt, sub_lt_iff_lt_add, mul_assoc] at h, + have hu₀ : 0 ≤ u - ⌊ξ⌋ * v, + { refine (zero_le_mul_right hv₁).mp ((lt_iff_add_one_le (-1 : ℤ) _).mp _), + replace h := h.1, + rw [← lt_sub_iff_add_lt, ← mul_assoc, ← sub_mul] at h, + exact_mod_cast h.trans_le ((mul_le_mul_right $ hv₀').mpr $ + (sub_le_sub_iff_left (u : ℝ)).mpr ((mul_le_mul_right hv₀).mpr (floor_le ξ))), }, + have hu₁ : u - ⌊ξ⌋ * v ≤ v, + { refine le_of_mul_le_mul_right (le_of_lt_add_one _) hv₁, + replace h := h.2, + rw [← sub_lt_iff_lt_add, ← mul_assoc, ← sub_mul, + ← add_lt_add_iff_right (v * (2 * v - 1) : ℝ), add_comm (1 : ℝ)] at h, + have := (mul_lt_mul_right $ hv₀').mpr ((sub_lt_sub_iff_left (u : ℝ)).mpr $ + (mul_lt_mul_right hv₀).mpr $ sub_right_lt_of_lt_add $ lt_floor_add_one ξ), + rw [sub_mul ξ, one_mul, ← sub_add, add_mul] at this, + exact_mod_cast this.trans h, }, + have huv_cop : is_coprime (u - ⌊ξ⌋ * v) v, + { rwa [sub_eq_add_neg, ← neg_mul, is_coprime.add_mul_right_left_iff], }, + refine ⟨lt_of_le_of_ne' hu₀ (λ hf, _), lt_of_le_of_ne hu₁ (λ hf, _)⟩; + { rw hf at huv_cop, + simp only [is_coprime_zero_left, is_coprime_self, is_unit_iff] at huv_cop, + cases huv_cop; linarith only [hv, huv_cop], }, +end + +-- The key step: the relevant inequality persists in the inductive step. +private +lemma aux₃ : |(fract ξ)⁻¹ - v / (u - ⌊ξ⌋ * v)| < ((u - ⌊ξ⌋ * v) * (2 * (u - ⌊ξ⌋ * v) - 1))⁻¹ := +begin + obtain ⟨hu₀, huv⟩ := aux₂ hv h, + have hξ₀ := aux₁ hv h, + set u' := u - ⌊ξ⌋ * v with hu', + have hu'ℝ : (u' : ℝ) = u - ⌊ξ⌋ * v := by exact_mod_cast hu', + rw ← hu'ℝ, + replace hu'ℝ := (eq_sub_iff_add_eq.mp hu'ℝ).symm, + obtain ⟨Hu, Hu'⟩ := aux₀ hu₀, + obtain ⟨Hv, Hv'⟩ := aux₀ (zero_lt_two.trans_le hv), + have H₁ := div_pos (div_pos Hv Hu) hξ₀, + replace h := h.2.2, + have h' : |fract ξ - u' / v| < (v * (2 * v - 1))⁻¹, + { rwa [hu'ℝ, add_div, mul_div_cancel _ Hv.ne', ← sub_sub, sub_right_comm] at h, }, + have H : (2 * u' - 1 : ℝ) ≤ (2 * v - 1) * fract ξ, + { replace h := (abs_lt.mp h).1, + have : (2 * (v : ℝ) - 1) * (-(v * (2 * v - 1))⁻¹ + u' / v) = 2 * u' - (1 + u') / v, + { field_simp [Hv.ne', Hv'.ne'], ring, }, + rw [hu'ℝ, add_div, mul_div_cancel _ Hv.ne', ← sub_sub, sub_right_comm, self_sub_floor, + lt_sub_iff_add_lt, ← mul_lt_mul_left Hv', this] at h, + refine has_le.le.trans _ h.le, + rw [sub_le_sub_iff_left, div_le_one Hv, add_comm], + exact_mod_cast huv, }, + have help₁ : ∀ {a b c : ℝ}, a ≠ 0 → b ≠ 0 → c ≠ 0 → + |a⁻¹ - b / c| = |(a - c / b) * (b / c / a)|, + { intros, rw abs_sub_comm, congr' 1, field_simp, ring }, + have help₂ : ∀ {a b c d : ℝ}, a ≠ 0 → b ≠ 0 → c ≠ 0 → d ≠ 0 → + (b * c)⁻¹ * (b / d / a) = (d * c * a)⁻¹, + { intros, field_simp, ring }, + calc + |(fract ξ)⁻¹ - v / u'| + = |(fract ξ - u' / v) * (v / u' / fract ξ)| : help₁ hξ₀.ne' Hv.ne' Hu.ne' + ... = |fract ξ - u' / v| * (v / u' / fract ξ) : by rw [abs_mul, abs_of_pos H₁, abs_sub_comm] + ... < (v * (2 * v - 1))⁻¹ * (v / u' / fract ξ) : (mul_lt_mul_right H₁).mpr h' + ... = (u' * (2 * v - 1) * fract ξ)⁻¹ : help₂ hξ₀.ne' Hv.ne' Hv'.ne' Hu.ne' + ... ≤ (u' * (2 * u' - 1))⁻¹ : by rwa [inv_le_inv (mul_pos (mul_pos Hu Hv') hξ₀) $ + mul_pos Hu Hu', mul_assoc, mul_le_mul_left Hu], +end + +-- The conditions `ass ξ u v` persist in the inductive step. +private lemma invariant : contfrac_legendre.ass (fract ξ)⁻¹ v (u - ⌊ξ⌋ * v) := +begin + refine ⟨_, λ huv, _, by exact_mod_cast aux₃ hv h⟩, + { rw [sub_eq_add_neg, ← neg_mul, is_coprime_comm, is_coprime.add_mul_right_left_iff], + exact h.1, }, + { obtain ⟨hv₀, hv₀'⟩ := aux₀ (zero_lt_two.trans_le hv), + have Hv : (v * (2 * v - 1) : ℝ)⁻¹ + v⁻¹ = 2 / (2 * v - 1), + { field_simp [hv₀.ne', hv₀'.ne'], ring, }, + have Huv : (u / v : ℝ) = ⌊ξ⌋ + v⁻¹, + { rw [sub_eq_iff_eq_add'.mp huv], field_simp [hv₀.ne'], }, + have h' := (abs_sub_lt_iff.mp h.2.2).1, + rw [Huv, ← sub_sub, sub_lt_iff_lt_add, self_sub_floor, Hv] at h', + rwa [lt_sub_iff_add_lt', (by ring : (v : ℝ) + -(1 / 2) = (2 * v - 1) / 2), + lt_inv (div_pos hv₀' zero_lt_two) (aux₁ hv h), inv_div], } +end + +omit h hv + +/-! +### The main result +-/ + +/-- The technical version of *Legendre's Theorem*. -/ +lemma exists_rat_eq_convergent' {v : ℕ} (h : contfrac_legendre.ass ξ u v) : + ∃ n, (u / v : ℚ) = ξ.convergent n := +begin + induction v using nat.strong_induction_on with v ih generalizing ξ u, + rcases lt_trichotomy v 1 with ht | rfl | ht, + { replace h := h.2.2, + simp only [nat.lt_one_iff.mp ht, nat.cast_zero, div_zero, tsub_zero, zero_mul, cast_zero, + inv_zero] at h, + exact false.elim (lt_irrefl _ $ (abs_nonneg ξ).trans_lt h), }, + { rw [nat.cast_one, div_one], + obtain ⟨_, h₁, h₂⟩ := h, + cases le_or_lt (u : ℝ) ξ with ht ht, + { use 0, + rw [convergent_zero, rat.coe_int_inj, eq_comm, floor_eq_iff], + convert and.intro ht (sub_lt_iff_lt_add'.mp (abs_lt.mp h₂).2); norm_num, }, + { replace h₁ := lt_sub_iff_add_lt'.mp (h₁ rfl), + have hξ₁ : ⌊ξ⌋ = u - 1, + { rw [floor_eq_iff, cast_sub, cast_one, sub_add_cancel], + exact ⟨(((sub_lt_sub_iff_left _).mpr one_half_lt_one).trans h₁).le, ht⟩, }, + cases eq_or_ne ξ ⌊ξ⌋ with Hξ Hξ, + { rw [Hξ, hξ₁, cast_sub, cast_one, ← sub_eq_add_neg, sub_lt_sub_iff_left] at h₁, + exact false.elim (lt_irrefl _ $ h₁.trans one_half_lt_one), }, + { have hξ₂ : ⌊(fract ξ)⁻¹⌋ = 1, + { rw [floor_eq_iff, cast_one, le_inv zero_lt_one (fract_pos.mpr Hξ), inv_one, + one_add_one_eq_two, inv_lt (fract_pos.mpr Hξ) zero_lt_two], + refine ⟨(fract_lt_one ξ).le, _⟩, + rw [fract, hξ₁, cast_sub, cast_one, lt_sub_iff_add_lt', sub_add], + convert h₁, + norm_num, }, + use 1, + simp [convergent, hξ₁, hξ₂, cast_sub, cast_one], } } }, + { obtain ⟨huv₀, huv₁⟩ := aux₂ (nat.cast_le.mpr ht) h, + have Hv : (v : ℚ) ≠ 0 := (nat.cast_pos.mpr (zero_lt_one.trans ht)).ne', + have huv₁' : (u - ⌊ξ⌋ * v).to_nat < v := by { zify, rwa to_nat_of_nonneg huv₀.le, }, + have inv : contfrac_legendre.ass (fract ξ)⁻¹ v (u - ⌊ξ⌋ * ↑v).to_nat := + (to_nat_of_nonneg huv₀.le).symm ▸ invariant (nat.cast_le.mpr ht) h, + obtain ⟨n, hn⟩ := ih (u - ⌊ξ⌋ * v).to_nat huv₁' inv, + use (n + 1), + rw [convergent_succ, ← hn, + (by exact_mod_cast to_nat_of_nonneg huv₀.le : ((u - ⌊ξ⌋ * v).to_nat : ℚ) = u - ⌊ξ⌋ * v), + ← coe_coe, inv_div, sub_div, mul_div_cancel _ Hv, add_sub_cancel'_right], } +end + +/-- The main result, *Legendre's Theorem* on rational approximation: +if `ξ` is a real number and `q` is a rational number such that `|ξ - q| < 1/(2*q.denom^2)`, +then `q` is a convergent of the continued fraction expansion of `ξ`. +This version uses `real.convergent`. -/ +lemma exists_rat_eq_convergent {q : ℚ} (h : |ξ - q| < 1 / (2 * q.denom ^ 2)) : + ∃ n, q = ξ.convergent n := +begin + refine q.num_div_denom ▸ exists_rat_eq_convergent' ⟨_, λ hd, _, _⟩, + { exact coprime_iff_nat_coprime.mpr (nat_abs_of_nat q.denom ▸ q.cop), }, + { rw ← q.denom_eq_one_iff.mp (nat.cast_eq_one.mp hd) at h, + simpa only [rat.coe_int_denom, nat.cast_one, one_pow, mul_one] using (abs_lt.mp h).1 }, + { obtain ⟨hq₀, hq₁⟩ := aux₀ (nat.cast_pos.mpr q.pos), + replace hq₁ := mul_pos hq₀ hq₁, + have hq₂ : (0 : ℝ) < 2 * (q.denom * q.denom) := mul_pos zero_lt_two (mul_pos hq₀ hq₀), + rw ← coe_coe at *, + rw [(by norm_cast : (q.num / q.denom : ℝ) = (q.num / q.denom : ℚ)), rat.num_div_denom], + exact h.trans + (by {rw [← one_div, sq, one_div_lt_one_div hq₂ hq₁, ← sub_pos], ring_nf, exact hq₀}), } +end + +/-- The main result, *Legendre's Theorem* on rational approximation: +if `ξ` is a real number and `q` is a rational number such that `|ξ - q| < 1/(2*q.denom^2)`, +then `q` is a convergent of the continued fraction expansion of `ξ`. +This is the version using `generalized_contined_fraction.convergents`. -/ +lemma exists_continued_fraction_convergent_eq_rat {q : ℚ} (h : |ξ - q| < 1 / (2 * q.denom ^ 2)) : + ∃ n, (generalized_continued_fraction.of ξ).convergents n = q := +begin + obtain ⟨n, hn⟩ := exists_rat_eq_convergent h, + exact ⟨n, hn.symm ▸ continued_fraction_convergent_eq_convergent ξ n⟩, +end + +end real From a2706b55e8d7f7e9b1f93143f0b88f2e34a11eea Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Wed, 10 May 2023 14:51:30 +0000 Subject: [PATCH 0985/1291] chore(*): add mathlib4 synchronization comments (#18977) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `category_theory.functor.flat` * `category_theory.preadditive.injective` * `category_theory.preadditive.injective_resolution` * `linear_algebra.projective_space.subspace` * `measure_theory.measure.ae_measurable` * `measure_theory.measure.measure_space` * `measure_theory.measure.sub` * `probability.probability_mass_function.basic` * `topology.partition_of_unity` --- src/category_theory/functor/flat.lean | 3 +++ src/category_theory/preadditive/injective.lean | 3 +++ src/category_theory/preadditive/injective_resolution.lean | 3 +++ src/linear_algebra/projective_space/subspace.lean | 3 +++ src/measure_theory/measure/ae_measurable.lean | 3 +++ src/measure_theory/measure/measure_space.lean | 3 +++ src/measure_theory/measure/sub.lean | 3 +++ src/probability/probability_mass_function/basic.lean | 3 +++ src/topology/partition_of_unity.lean | 3 +++ 9 files changed, 27 insertions(+) diff --git a/src/category_theory/functor/flat.lean b/src/category_theory/functor/flat.lean index 6eca7ce37093f..c03d26266d8e1 100644 --- a/src/category_theory/functor/flat.lean +++ b/src/category_theory/functor/flat.lean @@ -13,6 +13,9 @@ import category_theory.limits.shapes.finite_limits /-! # Representably flat functors +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define representably flat functors as functors such that the category of structured arrows over `X` is cofiltered for each `X`. This concept is also known as flat functors as in [Elephant] Remark C2.3.7, and this name is suggested by Mike Shulman in diff --git a/src/category_theory/preadditive/injective.lean b/src/category_theory/preadditive/injective.lean index d791d6ecc1294..bcb4693c4d9d9 100644 --- a/src/category_theory/preadditive/injective.lean +++ b/src/category_theory/preadditive/injective.lean @@ -8,6 +8,9 @@ import category_theory.preadditive.projective /-! # Injective objects and categories with enough injectives +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + An object `J` is injective iff every morphism into `J` can be obtained by extending a monomorphism. -/ diff --git a/src/category_theory/preadditive/injective_resolution.lean b/src/category_theory/preadditive/injective_resolution.lean index d002bb53a9b1e..e6e74beb9162a 100644 --- a/src/category_theory/preadditive/injective_resolution.lean +++ b/src/category_theory/preadditive/injective_resolution.lean @@ -9,6 +9,9 @@ import algebra.homology.single /-! # Injective resolutions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A injective resolution `I : InjectiveResolution Z` of an object `Z : C` consists of a `ℕ`-indexed cochain complex `I.cocomplex` of injective objects, along with a cochain map `I.ι` from cochain complex consisting just of `Z` in degree zero to `C`, diff --git a/src/linear_algebra/projective_space/subspace.lean b/src/linear_algebra/projective_space/subspace.lean index d15d8c751ec4a..aadabc860d9d3 100644 --- a/src/linear_algebra/projective_space/subspace.lean +++ b/src/linear_algebra/projective_space/subspace.lean @@ -9,6 +9,9 @@ import linear_algebra.projective_space.basic /-! # Subspaces of Projective Space +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define subspaces of a projective space, and show that the subspaces of a projective space form a complete lattice under inclusion. diff --git a/src/measure_theory/measure/ae_measurable.lean b/src/measure_theory/measure/ae_measurable.lean index 6860189ef0373..85d55ad2bf67a 100644 --- a/src/measure_theory/measure/ae_measurable.lean +++ b/src/measure_theory/measure/ae_measurable.lean @@ -8,6 +8,9 @@ import measure_theory.measure.measure_space /-! # Almost everywhere measurable functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A function is almost everywhere measurable if it coincides almost everywhere with a measurable function. This property, called `ae_measurable f μ`, is defined in the file `measure_space_def`. We discuss several of its properties that are analogous to properties of measurable functions. diff --git a/src/measure_theory/measure/measure_space.lean b/src/measure_theory/measure/measure_space.lean index 1fbaec96e712a..c19686b5b94cb 100644 --- a/src/measure_theory/measure/measure_space.lean +++ b/src/measure_theory/measure/measure_space.lean @@ -10,6 +10,9 @@ import topology.algebra.order.liminf_limsup /-! # Measure spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The definition of a measure and a measure space are in `measure_theory.measure_space_def`, with only a few basic properties. This file provides many more properties of these objects. This separation allows the measurability tactic to import only the file `measure_space_def`, and to diff --git a/src/measure_theory/measure/sub.lean b/src/measure_theory/measure/sub.lean index cb8d1b097292a..fb3ce98e4bde6 100644 --- a/src/measure_theory/measure/sub.lean +++ b/src/measure_theory/measure/sub.lean @@ -8,6 +8,9 @@ import measure_theory.measure.measure_space /-! # Subtraction of measures +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define `μ - ν` to be the least measure `τ` such that `μ ≤ τ + ν`. It is the equivalent of `(μ - ν) ⊔ 0` if `μ` and `ν` were signed measures. Compare with `ennreal.has_sub`. diff --git a/src/probability/probability_mass_function/basic.lean b/src/probability/probability_mass_function/basic.lean index b468119385902..b849e214b150b 100644 --- a/src/probability/probability_mass_function/basic.lean +++ b/src/probability/probability_mass_function/basic.lean @@ -9,6 +9,9 @@ import measure_theory.measure.measure_space /-! # Probability mass functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file is about probability mass functions or discrete probability measures: a function `α → ℝ≥0∞` such that the values have (infinite) sum `1`. diff --git a/src/topology/partition_of_unity.lean b/src/topology/partition_of_unity.lean index 866967b080840..a826173c07c43 100644 --- a/src/topology/partition_of_unity.lean +++ b/src/topology/partition_of_unity.lean @@ -13,6 +13,9 @@ import topology.urysohns_lemma /-! # Continuous partition of unity +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define `partition_of_unity (ι X : Type*) [topological_space X] (s : set X := univ)` to be a continuous partition of unity on `s` indexed by `ι`. More precisely, `f : partition_of_unity ι X s` is a collection of continuous functions `f i : C(X, ℝ)`, `i : ι`, such that From 28b2a92f2996d28e580450863c130955de0ed398 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Wed, 10 May 2023 20:34:53 +0000 Subject: [PATCH 0986/1291] feat(probability/kernel/measurable_integral): the integral against a kernel is strongly measurable (#18974) We also rename the measurability lemmas to use the same convention as in the file measure_theory/constructions/prod, which contains very similar lemmas for the integral against a measure (a particular case of what we are proving here). --- src/probability/kernel/composition.lean | 20 +- .../kernel/measurable_integral.lean | 269 ++++++++++++++---- src/probability/kernel/with_density.lean | 2 +- 3 files changed, 224 insertions(+), 67 deletions(-) diff --git a/src/probability/kernel/composition.lean b/src/probability/kernel/composition.lean index 625092688fc58..c77ed1f7c9dda 100644 --- a/src/probability/kernel/composition.lean +++ b/src/probability/kernel/composition.lean @@ -121,7 +121,7 @@ begin { intros i, have hm : measurable_set {p : (α × β) × γ | (p.1.2, p.2) ∈ f i}, from measurable_fst.snd.prod_mk measurable_snd (hf_meas i), - exact ((measurable_prod_mk_mem η hm).comp measurable_prod_mk_left).ae_measurable, }, + exact ((measurable_kernel_prod_mk_left hm).comp measurable_prod_mk_left).ae_measurable, }, end lemma comp_prod_fun_tsum_right (κ : kernel α β) (η : kernel (α × β) γ) [is_s_finite_kernel η] @@ -136,7 +136,7 @@ begin rw measure.sum_apply, exact measurable_prod_mk_left hs, }, rw [this, lintegral_tsum (λ n : ℕ, _)], - exact ((measurable_prod_mk_mem (seq η n) ((measurable_fst.snd.prod_mk measurable_snd) hs)).comp + exact ((measurable_kernel_prod_mk_left ((measurable_fst.snd.prod_mk measurable_snd) hs)).comp measurable_prod_mk_left).ae_measurable, end @@ -163,8 +163,8 @@ begin have hp_eq_mk : p = (p.fst, p.snd) := prod.mk.eta.symm, rw [hp_eq_mk, function.uncurry_apply_pair], }, rw this, - exact measurable_prod_mk_mem η (measurable_fst.snd.prod_mk measurable_snd hs), }, - exact measurable_lintegral κ h_meas, + exact measurable_kernel_prod_mk_left (measurable_fst.snd.prod_mk measurable_snd hs), }, + exact h_meas.lintegral_kernel_prod_right, end lemma measurable_comp_prod_fun (κ : kernel α β) [is_s_finite_kernel κ] @@ -181,8 +181,8 @@ begin have hp_eq_mk : p = (p.fst, p.snd) := prod.mk.eta.symm, rw [hp_eq_mk, function.uncurry_apply_pair], }, rw this, - exact measurable_prod_mk_mem (seq η n) (measurable_fst.snd.prod_mk measurable_snd hs), }, - exact measurable_lintegral κ h_meas, + exact measurable_kernel_prod_mk_left (measurable_fst.snd.prod_mk measurable_snd hs), }, + exact h_meas.lintegral_kernel_prod_right, end /-- Composition-Product of kernels. It verifies @@ -249,7 +249,7 @@ begin { ext1 ab, refl, }, rw this, refine measurable.comp _ measurable_prod_mk_left, - exact (measurable_lintegral η + exact (measurable.lintegral_kernel_prod_right ((simple_func.measurable _).comp (measurable_fst.snd.prod_mk measurable_snd))), }, rw lintegral_supr, rotate, @@ -262,8 +262,8 @@ begin simp only [simple_func.const_zero, simple_func.coe_piecewise, simple_func.coe_const, simple_func.coe_zero, set.piecewise_eq_indicator, lintegral_indicator_const hs], rw [comp_prod_apply κ η _ hs, ← lintegral_const_mul c _], - swap, { exact (measurable_prod_mk_mem η ((measurable_fst.snd.prod_mk measurable_snd) hs)).comp - measurable_prod_mk_left, }, + swap, { exact (measurable_kernel_prod_mk_left + ((measurable_fst.snd.prod_mk measurable_snd) hs)).comp measurable_prod_mk_left, }, congr, ext1 b, rw lintegral_indicator_const_comp measurable_prod_mk_left hs, @@ -658,7 +658,7 @@ lemma comp_assoc {δ : Type*} {mδ : measurable_space δ} (ξ : kernel γ δ) [i (ξ ∘ₖ η ∘ₖ κ) = ξ ∘ₖ (η ∘ₖ κ) := begin refine ext_fun (λ a f hf, _), - simp_rw [lintegral_comp _ _ _ hf, lintegral_comp _ _ _ (measurable_lintegral' ξ hf)], + simp_rw [lintegral_comp _ _ _ hf, lintegral_comp _ _ _ hf.lintegral_kernel], end lemma deterministic_comp_eq_map (hf : measurable f) (κ : kernel α β) [is_s_finite_kernel κ] : diff --git a/src/probability/kernel/measurable_integral.lean b/src/probability/kernel/measurable_integral.lean index a94cedbd09d20..94b0146398a24 100644 --- a/src/probability/kernel/measurable_integral.lean +++ b/src/probability/kernel/measurable_integral.lean @@ -8,75 +8,76 @@ import probability.kernel.basic /-! # Measurability of the integral against a kernel -The Lebesgue integral of a measurable function against a kernel is measurable. +The Lebesgue integral of a measurable function against a kernel is measurable. The Bochner integral +is strongly measurable. ## Main statements -* `probability_theory.kernel.measurable_lintegral`: the function `a ↦ ∫⁻ b, f a b ∂(κ a)` is - measurable, for an s-finite kernel `κ : kernel α β` and a function `f : α → β → ℝ≥0∞` such that - `function.uncurry f` is measurable. - +* `measurable.lintegral_kernel_prod_right`: the function `a ↦ ∫⁻ b, f a b ∂(κ a)` is measurable, + for an s-finite kernel `κ : kernel α β` and a function `f : α → β → ℝ≥0∞` such that `uncurry f` + is measurable. +* `measure_theory.strongly_measurable.integral_kernel_prod_right`: the function + `a ↦ ∫ b, f a b ∂(κ a)` is measurable, for an s-finite kernel `κ : kernel α β` and a function + `f : α → β → E` such that `uncurry f` is measurable. -/ -open measure_theory probability_theory - -open_locale measure_theory ennreal nnreal big_operators +open measure_theory probability_theory function set filter -namespace probability_theory.kernel +open_locale measure_theory ennreal topology -variables {α β ι : Type*} {mα : measurable_space α} {mβ : measurable_space β} +variables {α β γ : Type*} {mα : measurable_space α} {mβ : measurable_space β} + {mγ : measurable_space γ} + {κ : kernel α β} {η : kernel (α × β) γ} {a : α} -include mα mβ +namespace probability_theory +namespace kernel -/-- This is an auxiliary lemma for `measurable_prod_mk_mem`. -/ -lemma measurable_prod_mk_mem_of_finite (κ : kernel α β) {t : set (α × β)} (ht : measurable_set t) +/-- This is an auxiliary lemma for `measurable_kernel_prod_mk_left`. -/ +lemma measurable_kernel_prod_mk_left_of_finite {t : set (α × β)} (ht : measurable_set t) (hκs : ∀ a, is_finite_measure (κ a)) : - measurable (λ a, κ a {b | (a, b) ∈ t}) := + measurable (λ a, κ a (prod.mk a ⁻¹' t)) := begin -- `t` is a measurable set in the product `α × β`: we use that the product σ-algebra is generated -- by boxes to prove the result by induction. refine measurable_space.induction_on_inter generate_from_prod.symm is_pi_system_prod _ _ _ _ ht, { -- case `t = ∅` - simp only [set.mem_empty_iff_false, set.set_of_false, measure_empty, measurable_const], }, + simp only [preimage_empty, measure_empty, measurable_const], }, { -- case of a box: `t = t₁ ×ˢ t₂` for measurable sets `t₁` and `t₂` intros t' ht', simp only [set.mem_image2, set.mem_set_of_eq, exists_and_distrib_left] at ht', obtain ⟨t₁, ht₁, t₂, ht₂, rfl⟩ := ht', - simp only [set.prod_mk_mem_set_prod_eq], classical, - have h_eq_ite : (λ a, κ a {b : β | a ∈ t₁ ∧ b ∈ t₂}) = λ a, ite (a ∈ t₁) (κ a t₂) 0, + simp_rw mk_preimage_prod_right_eq_if, + have h_eq_ite : (λ a, κ a (ite (a ∈ t₁) t₂ ∅)) = λ a, ite (a ∈ t₁) (κ a t₂) 0, { ext1 a, split_ifs, - { simp only [h, true_and], refl, }, - { simp only [h, false_and, set.set_of_false, set.inter_empty, measure_empty], }, }, + exacts [rfl, measure_empty], }, rw h_eq_ite, exact measurable.ite ht₁ (kernel.measurable_coe κ ht₂) measurable_const }, { -- we assume that the result is true for `t` and we prove it for `tᶜ` intros t' ht' h_meas, - have h_eq_sdiff : ∀ a, {b : β | (a, b) ∈ t'ᶜ} = set.univ \ {b : β | (a, b) ∈ t'}, + have h_eq_sdiff : ∀ a, (prod.mk a ⁻¹' t'ᶜ) = set.univ \ (prod.mk a ⁻¹' t'), { intro a, ext1 b, - simp only [set.mem_compl_iff, set.mem_set_of_eq, set.mem_diff, set.mem_univ, true_and], }, + simp only [mem_compl_iff, mem_preimage, mem_diff, mem_univ, true_and], }, simp_rw h_eq_sdiff, - have : (λ a, κ a (set.univ \ {b : β | (a, b) ∈ t'})) - = (λ a, (κ a set.univ - κ a {b : β | (a, b) ∈ t'})), + have : (λ a, κ a (set.univ \ (prod.mk a ⁻¹' t'))) + = (λ a, (κ a set.univ - κ a (prod.mk a ⁻¹' t'))), { ext1 a, - rw [← set.diff_inter_self_eq_diff, set.inter_univ, measure_diff], - { exact set.subset_univ _, }, + rw [← set.diff_inter_self_eq_diff, set.inter_univ, measure_diff (set.subset_univ _)], { exact (@measurable_prod_mk_left α β _ _ a) t' ht', }, { exact measure_ne_top _ _, }, }, rw this, exact measurable.sub (kernel.measurable_coe κ measurable_set.univ) h_meas, }, { -- we assume that the result is true for a family of disjoint sets and prove it for their union intros f h_disj hf_meas hf, - have h_Union : (λ a, κ a {b : β | (a, b) ∈ ⋃ i, f i}) = λ a, κ a (⋃ i, {b : β | (a, b) ∈ f i}), + have h_Union : (λ a, κ a (prod.mk a ⁻¹' ⋃ i, f i)) = λ a, κ a (⋃ i, prod.mk a ⁻¹' f i), { ext1 a, congr' with b, - simp only [set.mem_Union, set.supr_eq_Union, set.mem_set_of_eq], - refl, }, + simp only [mem_Union, mem_preimage], }, rw h_Union, - have h_tsum : (λ a, κ a (⋃ i, {b : β | (a, b) ∈ f i})) = λ a, ∑' i, κ a {b : β | (a, b) ∈ f i}, + have h_tsum : (λ a, κ a (⋃ i, prod.mk a ⁻¹' f i)) = λ a, ∑' i, κ a (prod.mk a ⁻¹' f i), { ext1 a, rw measure_Union, { intros i j hij s hsi hsj b hbs, @@ -89,49 +90,74 @@ begin exact measurable.ennreal_tsum hf, }, end -lemma measurable_prod_mk_mem (κ : kernel α β) [is_s_finite_kernel κ] +lemma measurable_kernel_prod_mk_left [is_s_finite_kernel κ] {t : set (α × β)} (ht : measurable_set t) : - measurable (λ a, κ a {b | (a, b) ∈ t}) := + measurable (λ a, κ a (prod.mk a ⁻¹' t)) := begin rw ← kernel_sum_seq κ, - have : ∀ a, kernel.sum (seq κ) a {b : β | (a, b) ∈ t} = ∑' n, seq κ n a {b : β | (a, b) ∈ t}, + have : ∀ a, kernel.sum (seq κ) a (prod.mk a ⁻¹' t) = ∑' n, seq κ n a (prod.mk a ⁻¹' t), from λ a, kernel.sum_apply' _ _ (measurable_prod_mk_left ht), simp_rw this, refine measurable.ennreal_tsum (λ n, _), - exact measurable_prod_mk_mem_of_finite (seq κ n) ht infer_instance, + exact measurable_kernel_prod_mk_left_of_finite ht infer_instance, +end + +lemma measurable_kernel_prod_mk_left' [is_s_finite_kernel η] + {s : set (β × γ)} (hs : measurable_set s) (a : α) : + measurable (λ b, η (a, b) (prod.mk b ⁻¹' s)) := +begin + have : ∀ b, prod.mk b ⁻¹' s = {c | ((a, b), c) ∈ {p : (α × β) × γ | (p.1.2, p.2) ∈ s}}, + { intro b, refl, }, + simp_rw this, + refine (measurable_kernel_prod_mk_left _).comp measurable_prod_mk_left, + exact (measurable_fst.snd.prod_mk measurable_snd) hs, end -lemma measurable_lintegral_indicator_const (κ : kernel α β) [is_s_finite_kernel κ] - {t : set (α × β)} (ht : measurable_set t) (c : ℝ≥0∞) : +lemma measurable_kernel_prod_mk_right [is_s_finite_kernel κ] + {s : set (β × α)} (hs : measurable_set s) : + measurable (λ y, κ y ((λ x, (x, y)) ⁻¹' s)) := +measurable_kernel_prod_mk_left (measurable_set_swap_iff.mpr hs) + +end kernel + +open probability_theory.kernel + +section lintegral + +variables [is_s_finite_kernel κ] [is_s_finite_kernel η] + +/-- Auxiliary lemma for `measurable.lintegral_kernel_prod_right`. -/ +lemma kernel.measurable_lintegral_indicator_const {t : set (α × β)} (ht : measurable_set t) + (c : ℝ≥0∞) : measurable (λ a, ∫⁻ b, t.indicator (function.const (α × β) c) (a, b) ∂(κ a)) := begin simp_rw lintegral_indicator_const_comp measurable_prod_mk_left ht _, - exact measurable.const_mul (measurable_prod_mk_mem _ ht) c, + exact measurable.const_mul (measurable_kernel_prod_mk_left ht) c, end /-- For an s-finite kernel `κ` and a function `f : α → β → ℝ≥0∞` which is measurable when seen as a -map from `α × β` (hypothesis `measurable (function.uncurry f)`), the integral -`a ↦ ∫⁻ b, f a b ∂(κ a)` is measurable. -/ -theorem measurable_lintegral (κ : kernel α β) [is_s_finite_kernel κ] - {f : α → β → ℝ≥0∞} (hf : measurable (function.uncurry f)) : +map from `α × β` (hypothesis `measurable (uncurry f)`), the integral `a ↦ ∫⁻ b, f a b ∂(κ a)` is +measurable. -/ +lemma _root_.measurable.lintegral_kernel_prod_right {f : α → β → ℝ≥0∞} + (hf : measurable (uncurry f)) : measurable (λ a, ∫⁻ b, f a b ∂(κ a)) := begin - let F : ℕ → simple_func (α × β) ℝ≥0∞ := simple_func.eapprox (function.uncurry f), - have h : ∀ a, (⨆ n, F n a) = function.uncurry f a, - from simple_func.supr_eapprox_apply (function.uncurry f) hf, - simp only [prod.forall, function.uncurry_apply_pair] at h, + let F : ℕ → simple_func (α × β) ℝ≥0∞ := simple_func.eapprox (uncurry f), + have h : ∀ a, (⨆ n, F n a) = uncurry f a, + from simple_func.supr_eapprox_apply (uncurry f) hf, + simp only [prod.forall, uncurry_apply_pair] at h, simp_rw ← h, have : ∀ a, ∫⁻ b, (⨆ n, F n (a, b)) ∂(κ a) = ⨆ n, ∫⁻ b, F n (a, b) ∂(κ a), { intro a, rw lintegral_supr, { exact λ n, (F n).measurable.comp measurable_prod_mk_left, }, - { exact λ i j hij b, simple_func.monotone_eapprox (function.uncurry f) hij _, }, }, + { exact λ i j hij b, simple_func.monotone_eapprox (uncurry f) hij _, }, }, simp_rw this, refine measurable_supr (λ n, simple_func.induction _ _ (F n)), { intros c t ht, simp only [simple_func.const_zero, simple_func.coe_piecewise, simple_func.coe_const, simple_func.coe_zero, set.piecewise_eq_indicator], - exact measurable_lintegral_indicator_const κ ht c, }, + exact kernel.measurable_lintegral_indicator_const ht c, }, { intros g₁ g₂ h_disj hm₁ hm₂, simp only [simple_func.coe_add, pi.add_apply], have h_add : (λ a, ∫⁻ b, g₁ (a, b) + g₂ (a, b) ∂(κ a)) @@ -142,19 +168,150 @@ begin exact measurable.add hm₁ hm₂, }, end -lemma measurable_lintegral' (κ : kernel α β) [is_s_finite_kernel κ] - {f : β → ℝ≥0∞} (hf : measurable f) : - measurable (λ a, ∫⁻ b, f b ∂(κ a)) := -measurable_lintegral κ (hf.comp measurable_snd) +lemma _root_.measurable.lintegral_kernel_prod_right' {f : (α × β) → ℝ≥0∞} (hf : measurable f) : + measurable (λ a, ∫⁻ b, f (a, b) ∂(κ a)) := +begin + refine measurable.lintegral_kernel_prod_right _, + have : uncurry (λ (a : α) (b : β), f (a, b)) = f, + { ext x, rw [← @prod.mk.eta _ _ x, uncurry_apply_pair], }, + rwa this, +end + +lemma _root_.measurable.lintegral_kernel_prod_right'' {f : β × γ → ℝ≥0∞} (hf : measurable f) : + measurable (λ x, ∫⁻ y, f (x, y) ∂(η (a, x))) := +begin + change measurable ((λ x, ∫⁻ y, (λ u : (α × β) × γ, f (u.1.2, u.2)) (x, y) ∂(η x)) + ∘ (λ x, (a, x))), + refine (measurable.lintegral_kernel_prod_right' _).comp measurable_prod_mk_left, + exact hf.comp (measurable_fst.snd.prod_mk measurable_snd), +end -lemma measurable_set_lintegral (κ : kernel α β) [is_s_finite_kernel κ] - {f : α → β → ℝ≥0∞} (hf : measurable (function.uncurry f)) {s : set β} (hs : measurable_set s) : +lemma _root_.measurable.set_lintegral_kernel_prod_right + {f : α → β → ℝ≥0∞} (hf : measurable (uncurry f)) {s : set β} (hs : measurable_set s) : measurable (λ a, ∫⁻ b in s, f a b ∂(κ a)) := -by { simp_rw ← lintegral_restrict κ hs, exact measurable_lintegral _ hf } +by { simp_rw ← lintegral_restrict κ hs, exact hf.lintegral_kernel_prod_right } + +lemma _root_.measurable.lintegral_kernel_prod_left' {f : β × α → ℝ≥0∞} (hf : measurable f) : + measurable (λ y, ∫⁻ x, f (x, y) ∂(κ y)) := +(measurable_swap_iff.mpr hf).lintegral_kernel_prod_right' + +lemma _root_.measurable.lintegral_kernel_prod_left + {f : β → α → ℝ≥0∞} (hf : measurable (uncurry f)) : + measurable (λ y, ∫⁻ x, f x y ∂(κ y)) := +hf.lintegral_kernel_prod_left' + +lemma _root_.measurable.set_lintegral_kernel_prod_left + {f : β → α → ℝ≥0∞} (hf : measurable (uncurry f)) {s : set β} (hs : measurable_set s) : + measurable (λ b, ∫⁻ a in s, f a b ∂(κ b)) := +by { simp_rw ← lintegral_restrict κ hs, exact hf.lintegral_kernel_prod_left } + +lemma _root_.measurable.lintegral_kernel {f : β → ℝ≥0∞} (hf : measurable f) : + measurable (λ a, ∫⁻ b, f b ∂(κ a)) := +measurable.lintegral_kernel_prod_right (hf.comp measurable_snd) -lemma measurable_set_lintegral' (κ : kernel α β) [is_s_finite_kernel κ] +lemma _root_.measurable.set_lintegral_kernel {f : β → ℝ≥0∞} (hf : measurable f) {s : set β} (hs : measurable_set s) : measurable (λ a, ∫⁻ b in s, f b ∂(κ a)) := -measurable_set_lintegral κ (hf.comp measurable_snd) hs +measurable.set_lintegral_kernel_prod_right (hf.comp measurable_snd) hs + +end lintegral + +variables {E : Type*} [normed_add_comm_group E] [is_s_finite_kernel κ] [is_s_finite_kernel η] + +lemma measurable_set_kernel_integrable ⦃f : α → β → E⦄ (hf : strongly_measurable (uncurry f)) : + measurable_set {x | integrable (f x) (κ x)} := +begin + simp_rw [integrable, hf.of_uncurry_left.ae_strongly_measurable, true_and], + exact measurable_set_lt (measurable.lintegral_kernel_prod_right hf.ennnorm) measurable_const +end + +end probability_theory + +open probability_theory probability_theory.kernel + +namespace measure_theory + +variables {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] [complete_space E] + [is_s_finite_kernel κ] [is_s_finite_kernel η] + +lemma strongly_measurable.integral_kernel_prod_right + ⦃f : α → β → E⦄ (hf : strongly_measurable (uncurry f)) : + strongly_measurable (λ x, ∫ y, f x y ∂(κ x)) := +begin + classical, + borelize E, + haveI : topological_space.separable_space (range (uncurry f) ∪ {0} : set E) := + hf.separable_space_range_union_singleton, + let s : ℕ → simple_func (α × β) E := simple_func.approx_on _ hf.measurable + (range (uncurry f) ∪ {0}) 0 (by simp), + let s' : ℕ → α → simple_func β E := λ n x, (s n).comp (prod.mk x) measurable_prod_mk_left, + let f' : ℕ → α → E := λ n, {x | integrable (f x) (κ x)}.indicator + (λ x, (s' n x).integral (κ x)), + have hf' : ∀ n, strongly_measurable (f' n), + { intro n, refine strongly_measurable.indicator _ (measurable_set_kernel_integrable hf), + have : ∀ x, (s' n x).range.filter (λ x, x ≠ 0) ⊆ (s n).range, + { intros x, refine finset.subset.trans (finset.filter_subset _ _) _, intro y, + simp_rw [simple_func.mem_range], rintro ⟨z, rfl⟩, exact ⟨(x, z), rfl⟩ }, + simp only [simple_func.integral_eq_sum_of_subset (this _)], + refine finset.strongly_measurable_sum _ (λ x _, _), + refine (measurable.ennreal_to_real _).strongly_measurable.smul_const _, + simp only [simple_func.coe_comp, preimage_comp] {single_pass := tt}, + apply measurable_kernel_prod_mk_left, + exact (s n).measurable_set_fiber x }, + have h2f' : tendsto f' at_top (𝓝 (λ (x : α), ∫ (y : β), f x y ∂(κ x))), + { rw [tendsto_pi_nhds], intro x, + by_cases hfx : integrable (f x) (κ x), + { have : ∀ n, integrable (s' n x) (κ x), + { intro n, apply (hfx.norm.add hfx.norm).mono' (s' n x).ae_strongly_measurable, + apply eventually_of_forall, intro y, + simp_rw [s', simple_func.coe_comp], exact simple_func.norm_approx_on_zero_le _ _ (x, y) n }, + simp only [f', hfx, simple_func.integral_eq_integral _ (this _), indicator_of_mem, + mem_set_of_eq], + refine tendsto_integral_of_dominated_convergence (λ y, ‖f x y‖ + ‖f x y‖) + (λ n, (s' n x).ae_strongly_measurable) (hfx.norm.add hfx.norm) _ _, + { exact λ n, eventually_of_forall (λ y, simple_func.norm_approx_on_zero_le _ _ (x, y) n) }, + { refine eventually_of_forall (λ y, simple_func.tendsto_approx_on _ _ _), + apply subset_closure, + simp [-uncurry_apply_pair], } }, + { simp [f', hfx, integral_undef], } }, + exact strongly_measurable_of_tendsto _ hf' h2f', +end + +lemma strongly_measurable.integral_kernel_prod_right' + ⦃f : α × β → E⦄ (hf : strongly_measurable f) : + strongly_measurable (λ x, ∫ y, f (x, y) ∂(κ x)) := +by { rw [← uncurry_curry f] at hf, exact hf.integral_kernel_prod_right } + +lemma strongly_measurable.integral_kernel_prod_right'' + {f : β × γ → E} (hf : strongly_measurable f) : + strongly_measurable (λ x, ∫ y, f (x, y) ∂(η (a, x))) := +begin + change strongly_measurable ((λ x, ∫ y, (λ u : (α × β) × γ, f (u.1.2, u.2)) (x, y) ∂(η x)) + ∘ (λ x, (a, x))), + refine strongly_measurable.comp_measurable _ measurable_prod_mk_left, + refine measure_theory.strongly_measurable.integral_kernel_prod_right' _, + exact hf.comp_measurable (measurable_fst.snd.prod_mk measurable_snd), +end + +lemma strongly_measurable.integral_kernel_prod_left + ⦃f : β → α → E⦄ (hf : strongly_measurable (uncurry f)) : + strongly_measurable (λ y, ∫ x, f x y ∂(κ y)) := +(hf.comp_measurable measurable_swap).integral_kernel_prod_right' + +lemma strongly_measurable.integral_kernel_prod_left' + ⦃f : β × α → E⦄ (hf : strongly_measurable f) : + strongly_measurable (λ y, ∫ x, f (x, y) ∂(κ y)) := +(hf.comp_measurable measurable_swap).integral_kernel_prod_right' + +lemma strongly_measurable.integral_kernel_prod_left'' + {f : γ × β → E} (hf : strongly_measurable f) : + strongly_measurable (λ y, ∫ x, f (x, y) ∂(η (a, y))) := +begin + change strongly_measurable ((λ y, ∫ x, (λ u : γ × (α × β), f (u.1, u.2.2)) (x, y) ∂(η y)) + ∘ (λ x, (a, x))), + refine strongly_measurable.comp_measurable _ measurable_prod_mk_left, + refine measure_theory.strongly_measurable.integral_kernel_prod_left' _, + exact hf.comp_measurable (measurable_fst.prod_mk measurable_snd.snd), +end -end probability_theory.kernel +end measure_theory diff --git a/src/probability/kernel/with_density.lean b/src/probability/kernel/with_density.lean index e351b37fdaa48..5ab8e74f45cfb 100644 --- a/src/probability/kernel/with_density.lean +++ b/src/probability/kernel/with_density.lean @@ -51,7 +51,7 @@ def with_density (κ : kernel α β) [is_s_finite_kernel κ] (f : α → β → begin refine measure.measurable_of_measurable_coe _ (λ s hs, _), simp_rw with_density_apply _ hs, - exact measurable_set_lintegral κ hf hs, + exact hf.set_lintegral_kernel_prod_right hs, end, } : kernel α β)) (λ hf, 0) From 213b0cff7bc5ab6696ee07cceec80829ce42efec Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 10 May 2023 21:40:00 +0000 Subject: [PATCH 0987/1291] feat(algebra/order/to_interval_mod): add circular_order instance for add_circle (#17743) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This also provides us with the same instance for `real.angle`. Co-authored-by: Yaël Dillies Co-authored-by: Mauricio Collares --- src/algebra/order/to_interval_mod.lean | 205 +++++++++++++++++- .../trigonometric/angle.lean | 3 + src/order/circular.lean | 2 +- src/topology/instances/add_circle.lean | 3 + 4 files changed, 209 insertions(+), 4 deletions(-) diff --git a/src/algebra/order/to_interval_mod.lean b/src/algebra/order/to_interval_mod.lean index b547bcb2a117d..30cba2c693fb9 100644 --- a/src/algebra/order/to_interval_mod.lean +++ b/src/algebra/order/to_interval_mod.lean @@ -9,6 +9,7 @@ import algebra.order.archimedean import algebra.periodic import data.int.succ_pred import group_theory.quotient_group +import order.circular /-! # Reducing to an interval modulo its length @@ -32,8 +33,6 @@ interval. noncomputable theory -open add_comm_group - section linear_ordered_add_comm_group variables {α : Type*} [linear_ordered_add_comm_group α] [hα : archimedean α] {p : α} (hp : 0 < p) @@ -562,6 +561,8 @@ end end Ico_Ioc +open add_comm_group + lemma to_Ico_mod_eq_self : to_Ico_mod hp a b = b ↔ b ∈ set.Ico a (a + p) := by { rw [to_Ico_mod_eq_iff, and_iff_left], exact ⟨0, by simp⟩ } @@ -590,6 +591,37 @@ to_Ico_mod_add_right hp a lemma to_Ioc_mod_periodic (a : α) : function.periodic (to_Ioc_mod hp a) p := to_Ioc_mod_add_right hp a +-- helper lemmas for when `a = 0` +section zero + +lemma to_Ico_mod_zero_sub_comm (a b : α) : to_Ico_mod hp 0 (a - b) = p - to_Ioc_mod hp 0 (b - a) := +by rw [←neg_sub, to_Ico_mod_neg, neg_zero] + +lemma to_Ioc_mod_zero_sub_comm (a b : α) : to_Ioc_mod hp 0 (a - b) = p - to_Ico_mod hp 0 (b - a) := +by rw [←neg_sub, to_Ioc_mod_neg, neg_zero] + +lemma to_Ico_div_eq_sub (a b : α) : to_Ico_div hp a b = to_Ico_div hp 0 (b - a) := +by rw [to_Ico_div_sub_eq_to_Ico_div_add, zero_add] + +lemma to_Ioc_div_eq_sub (a b : α) : to_Ioc_div hp a b = to_Ioc_div hp 0 (b - a) := +by rw [to_Ioc_div_sub_eq_to_Ioc_div_add, zero_add] + +lemma to_Ico_mod_eq_sub (a b : α) : to_Ico_mod hp a b = to_Ico_mod hp 0 (b - a) + a := +by rw [to_Ico_mod_sub_eq_sub, zero_add, sub_add_cancel] + +lemma to_Ioc_mod_eq_sub (a b : α) : to_Ioc_mod hp a b = to_Ioc_mod hp 0 (b - a) + a := +by rw [to_Ioc_mod_sub_eq_sub, zero_add, sub_add_cancel] + +lemma to_Ico_mod_add_to_Ioc_mod_zero (a b : α) : + to_Ico_mod hp 0 (a - b) + to_Ioc_mod hp 0 (b - a) = p := +by rw [to_Ico_mod_zero_sub_comm, sub_add_cancel] + +lemma to_Ioc_mod_add_to_Ico_mod_zero (a b : α) : + to_Ioc_mod hp 0 (a - b) + to_Ico_mod hp 0 (b - a) = p := +by rw [add_comm, to_Ico_mod_add_to_Ioc_mod_zero] + +end zero + /-- `to_Ico_mod` as an equiv from the quotient. -/ @[simps symm_apply] def quotient_add_group.equiv_Ico_mod (a : α) : @@ -610,7 +642,12 @@ lemma quotient_add_group.equiv_Ico_mod_coe (a b : α) : quotient_add_group.equiv_Ico_mod hp a ↑b = ⟨to_Ico_mod hp a b, to_Ico_mod_mem_Ico hp a _⟩ := rfl -/-- `to_Ioc_mod` as an equiv from the quotient. -/ +@[simp] +lemma quotient_add_group.equiv_Ico_mod_zero (a : α) : + quotient_add_group.equiv_Ico_mod hp a 0 = ⟨to_Ico_mod hp a 0, to_Ico_mod_mem_Ico hp a _⟩ := +rfl + +/-- `to_Ioc_mod` as an equiv from the quotient. -/ @[simps symm_apply] def quotient_add_group.equiv_Ioc_mod (a : α) : (α ⧸ add_subgroup.zmultiples p) ≃ set.Ioc a (a + p) := @@ -630,8 +667,170 @@ lemma quotient_add_group.equiv_Ioc_mod_coe (a b : α) : quotient_add_group.equiv_Ioc_mod hp a ↑b = ⟨to_Ioc_mod hp a b, to_Ioc_mod_mem_Ioc hp a _⟩ := rfl +@[simp] +lemma quotient_add_group.equiv_Ioc_mod_zero (a : α) : + quotient_add_group.equiv_Ioc_mod hp a 0 = ⟨to_Ioc_mod hp a 0, to_Ioc_mod_mem_Ioc hp a _⟩ := +rfl + +/-! +### The circular order structure on `α ⧸ add_subgroup.zmultiples p` +-/ + +section circular + +private lemma to_Ixx_mod_iff (x₁ x₂ x₃ : α) : + to_Ico_mod hp x₁ x₂ ≤ to_Ioc_mod hp x₁ x₃ ↔ + to_Ico_mod hp 0 (x₂ - x₁) + to_Ico_mod hp 0 (x₁ - x₃) ≤ p := +by rw [to_Ico_mod_eq_sub, to_Ioc_mod_eq_sub _ x₁, add_le_add_iff_right, ←neg_sub x₁ x₃, + to_Ioc_mod_neg, neg_zero, le_sub_iff_add_le] + +private lemma to_Ixx_mod_cyclic_left {x₁ x₂ x₃ : α} + (h : to_Ico_mod hp x₁ x₂ ≤ to_Ioc_mod hp x₁ x₃) : + to_Ico_mod hp x₂ x₃ ≤ to_Ioc_mod hp x₂ x₁ := +begin + let x₂' := to_Ico_mod hp x₁ x₂, + let x₃' := to_Ico_mod hp x₂' x₃, + have h : x₂' ≤ to_Ioc_mod hp x₁ x₃' := by simpa, + have h₂₁ : x₂' < x₁ + p := to_Ico_mod_lt_right _ _ _, + have h₃₂ : x₃' - p < x₂' := sub_lt_iff_lt_add.2 (to_Ico_mod_lt_right _ _ _), + + suffices hequiv : x₃' ≤ to_Ioc_mod hp x₂' x₁, + { obtain ⟨z, hd⟩ : ∃ (z : ℤ), x₂ = x₂' + z • p := ((to_Ico_mod_eq_iff hp).1 rfl).2, + simpa [hd] }, + + cases le_or_lt x₃' (x₁ + p) with h₃₁ h₁₃, + { suffices hIoc₂₁ : to_Ioc_mod hp x₂' x₁ = x₁ + p, + { exact hIoc₂₁.symm.trans_ge h₃₁ }, + apply (to_Ioc_mod_eq_iff hp).2, + exact ⟨⟨h₂₁, by simp [left_le_to_Ico_mod]⟩, -1, by simp⟩ }, + + have hIoc₁₃ : to_Ioc_mod hp x₁ x₃' = x₃' - p, + { apply (to_Ioc_mod_eq_iff hp).2, + exact ⟨⟨lt_sub_iff_add_lt.2 h₁₃, le_of_lt (h₃₂.trans h₂₁)⟩, 1, by simp⟩ }, + have not_h₃₂ := (h.trans hIoc₁₃.le).not_lt, + contradiction +end + +private lemma to_Ixx_mod_antisymm (h₁₂₃ : to_Ico_mod hp a b ≤ to_Ioc_mod hp a c) + (h₁₃₂ : to_Ico_mod hp a c ≤ to_Ioc_mod hp a b) : + b ≡ a [PMOD p] ∨ c ≡ b [PMOD p] ∨ a ≡ c [PMOD p] := +begin + by_contra' h, + rw modeq_comm at h, + rw ←(not_modeq_iff_to_Ico_mod_eq_to_Ioc_mod hp).mp h.2.2 at h₁₂₃, + rw ←(not_modeq_iff_to_Ico_mod_eq_to_Ioc_mod hp).mp h.1 at h₁₃₂, + exact h.2.1 ((to_Ico_mod_inj _).1 $ h₁₃₂.antisymm h₁₂₃), +end + +private lemma to_Ixx_mod_total' (a b c : α) : + to_Ico_mod hp b a ≤ to_Ioc_mod hp b c ∨ to_Ico_mod hp b c ≤ to_Ioc_mod hp b a := +begin + /- an essential ingredient is the lemma sabing {a-b} + {b-a} = period if a ≠ b (and = 0 if a = b). + Thus if a ≠ b and b ≠ c then ({a-b} + {b-c}) + ({c-b} + {b-a}) = 2 * period, so one of + `{a-b} + {b-c}` and `{c-b} + {b-a}` must be `≤ period` -/ + have := congr_arg2 (+) + (to_Ico_mod_add_to_Ioc_mod_zero hp a b) (to_Ico_mod_add_to_Ioc_mod_zero hp c b), + rw [add_add_add_comm, add_comm (to_Ioc_mod _ _ _), add_add_add_comm, ←two_nsmul] at this, + replace := min_le_of_add_le_two_nsmul this.le, + rw min_le_iff at this, + rw [to_Ixx_mod_iff, to_Ixx_mod_iff], + refine this.imp (le_trans $ add_le_add_left _ _) (le_trans $ add_le_add_left _ _), + { apply to_Ico_mod_le_to_Ioc_mod }, + { apply to_Ico_mod_le_to_Ioc_mod } +end + +private lemma to_Ixx_mod_total (a b c : α) : + to_Ico_mod hp a b ≤ to_Ioc_mod hp a c ∨ to_Ico_mod hp c b ≤ to_Ioc_mod hp c a := +(to_Ixx_mod_total' _ _ _ _).imp_right $ to_Ixx_mod_cyclic_left _ + +private lemma to_Ixx_mod_trans {x₁ x₂ x₃ x₄ : α} + (h₁₂₃ : to_Ico_mod hp x₁ x₂ ≤ to_Ioc_mod hp x₁ x₃ + ∧ ¬to_Ico_mod hp x₃ x₂ ≤ to_Ioc_mod hp x₃ x₁) + (h₂₃₄ : to_Ico_mod hp x₂ x₄ ≤ to_Ioc_mod hp x₂ x₃ + ∧ ¬to_Ico_mod hp x₃ x₄ ≤ to_Ioc_mod hp x₃ x₂) : + to_Ico_mod hp x₁ x₄ ≤ to_Ioc_mod hp x₁ x₃ + ∧ ¬to_Ico_mod hp x₃ x₄ ≤ to_Ioc_mod hp x₃ x₁ := +begin + split, + { suffices h : ¬x₃ ≡ x₂ [PMOD p], + { have h₁₂₃' := to_Ixx_mod_cyclic_left _ (to_Ixx_mod_cyclic_left _ h₁₂₃.1), + have h₂₃₄' := to_Ixx_mod_cyclic_left _ (to_Ixx_mod_cyclic_left _ h₂₃₄.1), + rw [(not_modeq_iff_to_Ico_mod_eq_to_Ioc_mod hp).1 h] at h₂₃₄', + exact to_Ixx_mod_cyclic_left _ (h₁₂₃'.trans h₂₃₄') }, + by_contra, + rw [(modeq_iff_to_Ico_mod_eq_left hp).1 h] at h₁₂₃, + exact h₁₂₃.2 (left_lt_to_Ioc_mod _ _ _).le }, + { rw [not_le] at h₁₂₃ h₂₃₄ ⊢, + exact (h₁₂₃.2.trans_le (to_Ico_mod_le_to_Ioc_mod _ x₃ x₂)).trans h₂₃₄.2 }, +end + +namespace quotient_add_group +variables [hp' : fact (0 < p)] +include hp' + +instance : has_btw (α ⧸ add_subgroup.zmultiples p) := +{ btw := λ x₁ x₂ x₃, (equiv_Ico_mod hp'.out 0 (x₂ - x₁) : α) ≤ equiv_Ioc_mod hp'.out 0 (x₃ - x₁) } + +lemma btw_coe_iff' {x₁ x₂ x₃ : α} : + has_btw.btw (x₁ : α ⧸ add_subgroup.zmultiples p) x₂ x₃ ↔ + to_Ico_mod hp'.out 0 (x₂ - x₁) ≤ to_Ioc_mod hp'.out 0 (x₃ - x₁) := +iff.rfl + +-- maybe harder to use than the primed one? +lemma btw_coe_iff {x₁ x₂ x₃ : α} : + has_btw.btw (x₁ : α ⧸ add_subgroup.zmultiples p) x₂ x₃ ↔ + to_Ico_mod hp'.out x₁ x₂ ≤ to_Ioc_mod hp'.out x₁ x₃ := +by rw [btw_coe_iff', to_Ioc_mod_sub_eq_sub, to_Ico_mod_sub_eq_sub, zero_add, sub_le_sub_iff_right] + +instance circular_preorder : circular_preorder (α ⧸ add_subgroup.zmultiples p) := +{ btw_refl := λ x, show _ ≤ _, by simp [sub_self, hp'.out.le], + btw_cyclic_left := λ x₁ x₂ x₃ h, begin + induction x₁ using quotient_add_group.induction_on', + induction x₂ using quotient_add_group.induction_on', + induction x₃ using quotient_add_group.induction_on', + simp_rw [btw_coe_iff] at h ⊢, + apply to_Ixx_mod_cyclic_left _ h, + end, + sbtw := _, + sbtw_iff_btw_not_btw := λ _ _ _, iff.rfl, + sbtw_trans_left := λ x₁ x₂ x₃ x₄ (h₁₂₃ : _ ∧ _) (h₂₃₄ : _ ∧ _), show _ ∧ _, begin + induction x₁ using quotient_add_group.induction_on', + induction x₂ using quotient_add_group.induction_on', + induction x₃ using quotient_add_group.induction_on', + induction x₄ using quotient_add_group.induction_on', + simp_rw [btw_coe_iff] at h₁₂₃ h₂₃₄ ⊢, + apply to_Ixx_mod_trans _ h₁₂₃ h₂₃₄, + end } + +instance circular_order : circular_order (α ⧸ add_subgroup.zmultiples p) := +{ btw_antisymm := λ x₁ x₂ x₃ h₁₂₃ h₃₂₁, begin + induction x₁ using quotient_add_group.induction_on', + induction x₂ using quotient_add_group.induction_on', + induction x₃ using quotient_add_group.induction_on', + rw btw_cyclic at h₃₂₁, + simp_rw [btw_coe_iff] at h₁₂₃ h₃₂₁, + simp_rw ←modeq_iff_eq_mod_zmultiples, + exact to_Ixx_mod_antisymm _ h₁₂₃ h₃₂₁, + end, + btw_total := λ x₁ x₂ x₃, begin + induction x₁ using quotient_add_group.induction_on', + induction x₂ using quotient_add_group.induction_on', + induction x₃ using quotient_add_group.induction_on', + simp_rw [btw_coe_iff] at ⊢, + apply to_Ixx_mod_total, + end, + ..quotient_add_group.circular_preorder } + +end quotient_add_group + +end circular + end linear_ordered_add_comm_group +/-! +### Connections to `int.floor` and `int.fract` +-/ + section linear_ordered_field variables {α : Type*} [linear_ordered_field α] [floor_ring α] {p : α} (hp : 0 < p) diff --git a/src/analysis/special_functions/trigonometric/angle.lean b/src/analysis/special_functions/trigonometric/angle.lean index 958ca0e974bbb..422c30266f8b2 100644 --- a/src/analysis/special_functions/trigonometric/angle.lean +++ b/src/analysis/special_functions/trigonometric/angle.lean @@ -27,6 +27,9 @@ def angle : Type := add_circle (2 * π) namespace angle +instance : circular_order real.angle := +@add_circle.circular_order _ _ _ _ _ ⟨by norm_num [pi_pos]⟩ _ + @[continuity] lemma continuous_coe : continuous (coe : ℝ → angle) := continuous_quotient_mk diff --git a/src/order/circular.lean b/src/order/circular.lean index e768116353649..cd448697ca364 100644 --- a/src/order/circular.lean +++ b/src/order/circular.lean @@ -70,7 +70,7 @@ What is the correct generality of "rolling the necklace" open? At least, this wo `β × α` where `α` is a circular order and `β` is a linear order. What's next is to define circular groups and provide instances for `zmod n`, the usual circle group -`circle`, `real.angle`, and `roots_of_unity M`. What conditions do we need on `M` for this last one +`circle`, and `roots_of_unity M`. What conditions do we need on `M` for this last one to work? We should have circular order homomorphisms. The typical example is diff --git a/src/topology/instances/add_circle.lean b/src/topology/instances/add_circle.lean index 6847456e50aae..e84d83778331f 100644 --- a/src/topology/instances/add_circle.lean +++ b/src/topology/instances/add_circle.lean @@ -154,6 +154,9 @@ include hp variables (a : 𝕜) [archimedean 𝕜] +instance : circular_order (add_circle p) := +quotient_add_group.circular_order + /-- The equivalence between `add_circle p` and the half-open interval `[a, a + p)`, whose inverse is the natural quotient map. -/ def equiv_Ico : add_circle p ≃ Ico a (a + p) := quotient_add_group.equiv_Ico_mod hp.out a From 2d44d6823a96f9c79b7d1ab185918377be663424 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 10 May 2023 23:10:39 +0000 Subject: [PATCH 0988/1291] feat(data/finset/lattice): Distributivity lemmas (#18611) Dualise a few existing lemmas, protect `finset.sup_eq_bot_iff`/`finset.inf_eq_top_iff`, move `map_finset_sup`/`map_finset_inf` from `order.hom.lattice` to `data.finset.lattice`, make binders semi-implicit in `finset.disjoint_sup_left` and friends to avoid overly explicit binders in a local assumption after rewriting. Co-authored-by: Eric Wieser --- src/data/finset/lattice.lean | 82 +++++++++++++++++++++++++++++++----- 1 file changed, 71 insertions(+), 11 deletions(-) diff --git a/src/data/finset/lattice.lean b/src/data/finset/lattice.lean index 285a898550334..208cbac617b55 100644 --- a/src/data/finset/lattice.lean +++ b/src/data/finset/lattice.lean @@ -5,6 +5,7 @@ Authors: Mario Carneiro -/ import data.finset.fold import data.finset.option +import data.finset.pi import data.finset.prod import data.multiset.lattice import order.complete_lattice @@ -17,7 +18,7 @@ import order.hom.lattice > Any changes to this file require a corresponding PR to mathlib4. -/ -variables {F α β γ ι : Type*} +variables {F α β γ ι κ : Type*} namespace finset open multiset order_dual @@ -226,8 +227,7 @@ lemma sup_mem @sup_induction _ _ _ _ _ _ (∈ s) w₁ w₂ h @[simp] -lemma sup_eq_bot_iff (f : β → α) - (S : finset β) : S.sup f = ⊥ ↔ ∀ s ∈ S, f s = ⊥ := +protected lemma sup_eq_bot_iff (f : β → α) (S : finset β) : S.sup f = ⊥ ↔ ∀ s ∈ S, f s = ⊥ := begin classical, induction S using finset.induction with a S haS hi; @@ -377,8 +377,7 @@ lemma inf_mem @inf_induction _ _ _ _ _ _ (∈ s) w₁ w₂ h @[simp] -lemma inf_eq_top_iff (f : β → α) - (S : finset β) : S.inf f = ⊤ ↔ ∀ s ∈ S, f s = ⊤ := +protected lemma inf_eq_top_iff (f : β → α) (S : finset β) : S.inf f = ⊤ ↔ ∀ s ∈ S, f s = ⊤ := @finset.sup_eq_bot_iff αᵒᵈ _ _ _ _ _ end inf @@ -396,7 +395,7 @@ section distrib_lattice variables [distrib_lattice α] section order_bot -variables [order_bot α] {s : finset β} {f : β → α} {a : α} +variables [order_bot α] {s : finset ι} {t : finset κ} {f : ι → α} {g : κ → α} {a : α} lemma sup_inf_distrib_left (s : finset ι) (f : ι → α) (a : α) : a ⊓ s.sup f = s.sup (λ i, a ⊓ f i) := @@ -410,16 +409,20 @@ lemma sup_inf_distrib_right (s : finset ι) (f : ι → α) (a : α) : s.sup f ⊓ a = s.sup (λ i, f i ⊓ a) := by { rw [_root_.inf_comm, s.sup_inf_distrib_left], simp_rw _root_.inf_comm } -protected lemma disjoint_sup_right : disjoint a (s.sup f) ↔ ∀ i ∈ s, disjoint a (f i) := -by simp only [disjoint_iff, sup_inf_distrib_left, sup_eq_bot_iff] +protected lemma disjoint_sup_right : disjoint a (s.sup f) ↔ ∀ ⦃i⦄, i ∈ s → disjoint a (f i) := +by simp only [disjoint_iff, sup_inf_distrib_left, finset.sup_eq_bot_iff] -protected lemma disjoint_sup_left : disjoint (s.sup f) a ↔ ∀ i ∈ s, disjoint (f i) a := -by simp only [disjoint_iff, sup_inf_distrib_right, sup_eq_bot_iff] +protected lemma disjoint_sup_left : disjoint (s.sup f) a ↔ ∀ ⦃i⦄, i ∈ s → disjoint (f i) a := +by simp only [disjoint_iff, sup_inf_distrib_right, finset.sup_eq_bot_iff] + +lemma sup_inf_sup (s : finset ι) (t : finset κ) (f : ι → α) (g : κ → α) : + s.sup f ⊓ t.sup g = (s ×ˢ t).sup (λ i, f i.1 ⊓ g i.2) := +by simp_rw [finset.sup_inf_distrib_right, finset.sup_inf_distrib_left, sup_product_left] end order_bot section order_top -variables [order_top α] +variables [order_top α] {f : ι → α} {g : κ → α} {s : finset ι} {t : finset κ} {a : α} lemma inf_sup_distrib_left (s : finset ι) (f : ι → α) (a : α) : a ⊔ s.inf f = s.inf (λ i, a ⊔ f i) := @@ -429,7 +432,49 @@ lemma inf_sup_distrib_right (s : finset ι) (f : ι → α) (a : α) : s.inf f ⊔ a = s.inf (λ i, f i ⊔ a) := @sup_inf_distrib_right αᵒᵈ _ _ _ _ _ _ +protected lemma codisjoint_inf_right : codisjoint a (s.inf f) ↔ ∀ ⦃i⦄, i ∈ s → codisjoint a (f i) := +@finset.disjoint_sup_right αᵒᵈ _ _ _ _ _ _ + +protected lemma codisjoint_inf_left : codisjoint (s.inf f) a ↔ ∀ ⦃i⦄, i ∈ s → codisjoint (f i) a := +@finset.disjoint_sup_left αᵒᵈ _ _ _ _ _ _ + +lemma inf_sup_inf (s : finset ι) (t : finset κ) (f : ι → α) (g : κ → α) : + s.inf f ⊔ t.inf g = (s ×ˢ t).inf (λ i, f i.1 ⊔ g i.2) := +@sup_inf_sup αᵒᵈ _ _ _ _ _ _ _ _ + end order_top + +section bounded_order +variables [bounded_order α] [decidable_eq ι] + +--TODO: Extract out the obvious isomorphism `(insert i s).pi t ≃ t i ×ˢ s.pi t` from this proof +lemma inf_sup {κ : ι → Type*} (s : finset ι) (t : Π i, finset (κ i)) (f : Π i, κ i → α) : + s.inf (λ i, (t i).sup (f i)) = (s.pi t).sup (λ g, s.attach.inf $ λ i, f _ $ g _ i.prop) := +begin + induction s using finset.induction with i s hi ih, + { simp }, + rw [inf_insert, ih, attach_insert, sup_inf_sup], + refine eq_of_forall_ge_iff (λ c, _), + simp only [subtype.val_eq_coe, finset.sup_le_iff, mem_product, mem_pi, and_imp, prod.forall, + inf_insert, inf_image], + refine ⟨λ h g hg, h (g i $ mem_insert_self _ _) (λ j hj, g j $ mem_insert_of_mem hj) + (hg _ $ mem_insert_self _ _) (λ j hj, hg _ $ mem_insert_of_mem hj), λ h a g ha hg, _⟩, + -- TODO: This `have` must be named to prevent it being shadowed by the internal `this` in `simpa` + have aux : ∀ j : {x // x ∈ s}, ↑j ≠ i := λ j : s, ne_of_mem_of_not_mem j.2 hi, + simpa only [cast_eq, dif_pos, function.comp, subtype.coe_mk, dif_neg, aux] + using h (λ j hj, if hji : j = i then cast (congr_arg κ hji.symm) a + else g _ $ mem_of_mem_insert_of_ne hj hji) _, + simp_rw mem_insert, + rintro j (rfl | hj), + { simpa }, + { simpa [ne_of_mem_of_not_mem hj hi] using hg _ _ } +end + +lemma sup_inf {κ : ι → Type*} (s : finset ι) (t : Π i, finset (κ i)) (f : Π i, κ i → α) : + s.sup (λ i, (t i).inf (f i)) = (s.pi t).inf (λ g, s.attach.sup $ λ i, f _ $ g _ i.2) := +@inf_sup αᵒᵈ _ _ _ _ _ _ _ _ + +end bounded_order end distrib_lattice section boolean_algebra @@ -456,6 +501,21 @@ begin { rw [inf_cons, inf_cons, h, inf_sdiff] } end +lemma inf_himp_right (s : finset ι) (f : ι → α) (a : α) : s.inf (λ b, f b ⇨ a) = s.sup f ⇨ a := +@sup_sdiff_left αᵒᵈ _ _ _ _ _ + +lemma sup_himp_right (hs : s.nonempty) (f : ι → α) (a : α) : s.sup (λ b, f b ⇨ a) = s.inf f ⇨ a := +@inf_sdiff_left αᵒᵈ _ _ _ hs _ _ + +lemma sup_himp_left (hs : s.nonempty) (f : ι → α) (a : α) : s.sup (λ b, a ⇨ f b) = a ⇨ s.sup f := +@inf_sdiff_right αᵒᵈ _ _ _ hs _ _ + +@[simp] protected lemma compl_sup (s : finset ι) (f : ι → α) : (s.sup f)ᶜ = s.inf (λ i, (f i)ᶜ) := +map_finset_sup (order_iso.compl α) _ _ + +@[simp] protected lemma compl_inf (s : finset ι) (f : ι → α) : (s.inf f)ᶜ = s.sup (λ i, (f i)ᶜ) := +map_finset_inf (order_iso.compl α) _ _ + end boolean_algebra section linear_order From ec4b2eeb50364487f80421c0b4c41328a611f30d Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 10 May 2023 23:10:40 +0000 Subject: [PATCH 0989/1291] chore(topology/instances/ennreal): missing smul lemmas (#18980) --- src/data/real/ennreal.lean | 9 +++++++++ src/measure_theory/measure/outer_measure.lean | 14 ++++++++++++-- src/topology/instances/ennreal.lean | 14 ++++++++++++++ 3 files changed, 35 insertions(+), 2 deletions(-) diff --git a/src/data/real/ennreal.lean b/src/data/real/ennreal.lean index dabfc510a17e6..fda281f3e2bfe 100644 --- a/src/data/real/ennreal.lean +++ b/src/data/real/ennreal.lean @@ -312,6 +312,7 @@ def of_nnreal_hom : ℝ≥0 →+* ℝ≥0∞ := @[simp] lemma coe_of_nnreal_hom : ⇑of_nnreal_hom = coe := rfl +-- TODO: generalize some of these (and subsequent lemmas about `smul`) to `with_top α` section actions /-- A `mul_action` over `ℝ≥0∞` restricts to a `mul_action` over `ℝ≥0`. -/ @@ -388,6 +389,14 @@ begin split_ifs, { simp [h] }, { exact with_top.top_mul h } end @[simp] lemma top_mul_top : ∞ * ∞ = ∞ := with_top.top_mul_top +lemma smul_top {R} [has_zero R] [smul_with_zero R ℝ≥0∞] [is_scalar_tower R ℝ≥0∞ ℝ≥0∞] + [no_zero_smul_divisors R ℝ≥0∞] (c : R) : + c • ∞ = (if c = 0 then 0 else ∞) := +begin + rw [←smul_one_mul, mul_top], + simp_rw [smul_eq_zero, or_iff_left one_ne_zero], +end + lemma top_pow {n:ℕ} (h : 0 < n) : ∞^n = ∞ := nat.le_induction (pow_one _) (λ m hm hm', by rw [pow_succ, hm', top_mul_top]) _ (nat.succ_le_of_lt h) diff --git a/src/measure_theory/measure/outer_measure.lean b/src/measure_theory/measure/outer_measure.lean index e49e132c70227..af24e0abc4afe 100644 --- a/src/measure_theory/measure/outer_measure.lean +++ b/src/measure_theory/measure/outer_measure.lean @@ -346,8 +346,7 @@ by have := supr_apply (λ b, cond b m₁ m₂) s; theorem smul_supr [has_smul R ℝ≥0∞] [is_scalar_tower R ℝ≥0∞ ℝ≥0∞] {ι} (f : ι → outer_measure α) (c : R) : c • (⨆ i, f i) = ⨆ i, c • f i := -ext $ λ s, by simp only [smul_apply, supr_apply, ←smul_one_mul c (f _ _), - ←smul_one_mul c (supr _), ennreal.mul_supr] +ext $ λ s, by simp only [smul_apply, supr_apply, ennreal.smul_supr] end supremum @@ -1069,6 +1068,17 @@ variables (m : Π (s : α), P s → ℝ≥0∞) to all objects by defining it to be `∞` on the objects not in the class. -/ def extend (s : α) : ℝ≥0∞ := ⨅ h : P s, m s h +lemma smul_extend {R} [has_zero R] [smul_with_zero R ℝ≥0∞] [is_scalar_tower R ℝ≥0∞ ℝ≥0∞] + [no_zero_smul_divisors R ℝ≥0∞] {c : R} (hc : c ≠ 0) : + c • extend m = extend (λ s h, c • m s h) := +begin + ext1 s, + dsimp [extend], + by_cases h : P s, + { simp [h] }, + { simp [h, ennreal.smul_top, hc] }, +end + lemma extend_eq {s : α} (h : P s) : extend m s = m s h := by simp [extend, h] diff --git a/src/topology/instances/ennreal.lean b/src/topology/instances/ennreal.lean index 4937ad6818752..0fb4fb73cec59 100644 --- a/src/topology/instances/ennreal.lean +++ b/src/topology/instances/ennreal.lean @@ -612,6 +612,16 @@ by simp only [Sup_eq_supr, mul_supr] lemma supr_mul {ι : Sort*} {f : ι → ℝ≥0∞} {a : ℝ≥0∞} : supr f * a = ⨆i, f i * a := by rw [mul_comm, mul_supr]; congr; funext; rw [mul_comm] +lemma smul_supr {ι : Sort*} {R} [has_smul R ℝ≥0∞] [is_scalar_tower R ℝ≥0∞ ℝ≥0∞] + (f : ι → ℝ≥0∞) (c : R) : + c • (⨆ i, f i) = ⨆ i, c • f i := +by simp only [←smul_one_mul c (f _), ←smul_one_mul c (supr _), ennreal.mul_supr] + +lemma smul_Sup {R} [has_smul R ℝ≥0∞] [is_scalar_tower R ℝ≥0∞ ℝ≥0∞] + (s : set ℝ≥0∞) (c : R) : + c • Sup s = ⨆ i ∈ s, c • i := +by simp_rw [←smul_one_mul c (Sup _), ennreal.mul_Sup, smul_one_mul] + lemma supr_div {ι : Sort*} {f : ι → ℝ≥0∞} {a : ℝ≥0∞} : supr f / a = ⨆i, f i / a := supr_mul @@ -857,6 +867,10 @@ has_sum.tsum_eq this protected lemma tsum_mul_right : (∑'i, f i * a) = (∑'i, f i) * a := by simp [mul_comm, ennreal.tsum_mul_left] +protected lemma tsum_const_smul {R} [has_smul R ℝ≥0∞] [is_scalar_tower R ℝ≥0∞ ℝ≥0∞] (a : R) : + ∑' i, a • f i = a • ∑' i, f i := +by simpa only [smul_one_mul] using @ennreal.tsum_mul_left _ (a • 1) _ + @[simp] lemma tsum_supr_eq {α : Type*} (a : α) {f : α → ℝ≥0∞} : ∑'b:α, (⨆ (h : a = b), f b) = f a := le_antisymm From 5dc275ec639221ca4d5f56938eb966f6ad9bc89f Mon Sep 17 00:00:00 2001 From: Xavier Roblot <46200072+xroblot@users.noreply.github.com> Date: Thu, 11 May 2023 05:57:35 +0000 Subject: [PATCH 0990/1291] feat(analysis/normed/order/lattice): add has_solid_norm (#18554) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit See [this Zulip thread](https://leanprover.zulipchat.com/#narrow/stream/116395-maths/topic/Solid) Co-authored-by: Yaël Dillies --- src/algebra/order/lattice_group.lean | 25 +++++++++++++ src/analysis/normed/order/lattice.lean | 36 ++++++++++++++----- src/measure_theory/function/lp_order.lean | 2 +- src/topology/continuous_function/bounded.lean | 2 +- 4 files changed, 54 insertions(+), 11 deletions(-) diff --git a/src/algebra/order/lattice_group.lean b/src/algebra/order/lattice_group.lean index 0dbc871a56fc8..47482f0e26494 100644 --- a/src/algebra/order/lattice_group.lean +++ b/src/algebra/order/lattice_group.lean @@ -7,6 +7,8 @@ import algebra.group_power.basic -- Needed for squares import algebra.order.group.abs import tactic.nth_rewrite +import order.closure + /-! # Lattice ordered groups @@ -454,3 +456,26 @@ begin end end lattice_ordered_comm_group + +namespace lattice_ordered_add_comm_group + +variables {β : Type u} [lattice β] [add_comm_group β] + +section solid + +/-- A subset `s ⊆ β`, with `β` an `add_comm_group` with a `lattice` structure, is solid if for +all `x ∈ s` and all `y ∈ β` such that `|y| ≤ |x|`, then `y ∈ s`. -/ +def is_solid (s : set β) : Prop := ∀ ⦃x⦄, x ∈ s → ∀ ⦃y⦄, |y| ≤ |x| → y ∈ s + +/-- The solid closure of a subset `s` is the smallest superset of `s` that is solid. -/ +def solid_closure (s : set β) : set β := {y | ∃ x ∈ s, |y| ≤ |x|} + +lemma is_solid_solid_closure (s : set β) : is_solid (solid_closure s) := +λ x ⟨y, hy, hxy⟩ z hzx, ⟨y, hy, hzx.trans hxy⟩ + +lemma solid_closure_min (s t : set β) (h1 : s ⊆ t) (h2 : is_solid t) : solid_closure s ⊆ t := +λ _ ⟨_, hy, hxy⟩, h2 (h1 hy) hxy + +end solid + +end lattice_ordered_add_comm_group diff --git a/src/analysis/normed/order/lattice.lean b/src/analysis/normed/order/lattice.lean index 6ff7502e02a51..4370935d7c666 100644 --- a/src/analysis/normed/order/lattice.lean +++ b/src/analysis/normed/order/lattice.lean @@ -31,13 +31,36 @@ normed, lattice, ordered, group -/ /-! -### Normed lattice orderd groups +### Normed lattice ordered groups Motivated by the theory of Banach Lattices, this section introduces normed lattice ordered groups. -/ local notation (name := abs) `|`a`|` := abs a +section solid_norm + +/-- Let `α` be an `add_comm_group` with a `lattice` structure. A norm on `α` is *solid* if, for `a` +and `b` in `α`, with absolute values `|a|` and `|b|` respectively, `|a| ≤ |b|` implies `‖a‖ ≤ ‖b‖`. +-/ +class has_solid_norm (α : Type*) [normed_add_comm_group α] [lattice α] : Prop := +(solid : ∀ ⦃x y : α⦄, |x| ≤ |y| → ‖x‖ ≤ ‖y‖) + +variables {α : Type*} [normed_add_comm_group α] [lattice α] [has_solid_norm α] + +lemma norm_le_norm_of_abs_le_abs {a b : α} (h : |a| ≤ |b|) : ‖a‖ ≤ ‖b‖ := has_solid_norm.solid h + +/-- If `α` has a solid norm, then the balls centered at the origin of `α` are solid sets. -/ +lemma lattice_ordered_add_comm_group.is_solid_ball (r : ℝ) : + lattice_ordered_add_comm_group.is_solid (metric.ball (0 : α) r) := +λ _ hx _ hxy, mem_ball_zero_iff.mpr ((has_solid_norm.solid hxy).trans_lt (mem_ball_zero_iff.mp hx)) + +instance : has_solid_norm ℝ := ⟨λ _ _, id⟩ + +instance : has_solid_norm ℚ := ⟨λ _ _ _, by simpa only [norm, ← rat.cast_abs, rat.cast_le]⟩ + +end solid_norm + /-- Let `α` be a normed commutative group equipped with a partial order covariant with addition, with respect which `α` forms a lattice. Suppose that `α` is *solid*, that is to say, for `a` and `b` in @@ -45,16 +68,11 @@ respect which `α` forms a lattice. Suppose that `α` is *solid*, that is to say said to be a normed lattice ordered group. -/ class normed_lattice_add_comm_group (α : Type*) - extends normed_add_comm_group α, lattice α := + extends normed_add_comm_group α, lattice α, has_solid_norm α := (add_le_add_left : ∀ a b : α, a ≤ b → ∀ c : α, c + a ≤ c + b) -(solid : ∀ a b : α, |a| ≤ |b| → ‖a‖ ≤ ‖b‖) - -lemma solid {α : Type*} [normed_lattice_add_comm_group α] {a b : α} (h : |a| ≤ |b|) : ‖a‖ ≤ ‖b‖ := -normed_lattice_add_comm_group.solid a b h instance : normed_lattice_add_comm_group ℝ := -{ add_le_add_left := λ _ _ h _, add_le_add le_rfl h, - solid := λ _ _, id, } +{ add_le_add_left := λ _ _ h _, add_le_add le_rfl h,} /-- A normed lattice ordered group is an ordered additive commutative group @@ -64,7 +82,7 @@ instance normed_lattice_add_comm_group_to_ordered_add_comm_group {α : Type*} [h : normed_lattice_add_comm_group α] : ordered_add_comm_group α := { ..h } variables {α : Type*} [normed_lattice_add_comm_group α] -open lattice_ordered_comm_group +open lattice_ordered_comm_group has_solid_norm lemma dual_solid (a b : α) (h: b⊓-b ≤ a⊓-a) : ‖a‖ ≤ ‖b‖ := begin diff --git a/src/measure_theory/function/lp_order.lean b/src/measure_theory/function/lp_order.lean index 88775c495f921..fa528399bfcf1 100644 --- a/src/measure_theory/function/lp_order.lean +++ b/src/measure_theory/function/lp_order.lean @@ -94,7 +94,7 @@ instance [fact (1 ≤ p)] : normed_lattice_add_comm_group (Lp E p μ) := refine snorm_mono_ae _, filter_upwards [hfg, Lp.coe_fn_abs f, Lp.coe_fn_abs g] with x hx hxf hxg, rw [hxf, hxg] at hx, - exact solid hx, + exact has_solid_norm.solid hx, end, ..Lp.lattice, ..Lp.normed_add_comm_group, } diff --git a/src/topology/continuous_function/bounded.lean b/src/topology/continuous_function/bounded.lean index 32146f3e3f16c..7207cb70e936d 100644 --- a/src/topology/continuous_function/bounded.lean +++ b/src/topology/continuous_function/bounded.lean @@ -1328,7 +1328,7 @@ instance : normed_lattice_add_comm_group (α →ᵇ β) := solid := begin intros f g h, - have i1: ∀ t, ‖f t‖ ≤ ‖g t‖ := λ t, solid (h t), + have i1: ∀ t, ‖f t‖ ≤ ‖g t‖ := λ t, has_solid_norm.solid (h t), rw norm_le (norm_nonneg _), exact λ t, (i1 t).trans (norm_coe_le_norm g t), end, From 781cb2eed038c4caf53bdbd8d20a95e5822d77df Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Thu, 11 May 2023 07:22:30 +0000 Subject: [PATCH 0991/1291] chore(*): add mathlib4 synchronization comments (#18984) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.category.Group.epi_mono` * `algebra.category.Group.equivalence_Group_AddGroup` * `algebra.continued_fractions.computation.correctness_terminating` * `analysis.convex.partition_of_unity` * `analysis.normed.group.controlled_closure` * `analysis.normed_space.compact_operator` * `category_theory.abelian.pseudoelements` * `category_theory.preadditive.hom_orthogonal` * `category_theory.sites.cover_preserving` * `computability.halting` * `computability.partrec_code` * `dynamics.ergodic.ergodic` * `dynamics.ergodic.measure_preserving` * `measure_theory.covering.vitali_family` * `measure_theory.decomposition.unsigned_hahn` * `measure_theory.group.arithmetic` * `measure_theory.lattice` * `measure_theory.measure.mutually_singular` * `measure_theory.measure.open_pos` * `probability.conditional_probability` * `topology.algebra.equicontinuity` * `topology.category.Profinite.as_limit` * `topology.category.Top.limits.products` * `topology.continuous_function.locally_constant` --- src/algebra/category/Group/epi_mono.lean | 3 +++ src/algebra/category/Group/equivalence_Group_AddGroup.lean | 3 +++ .../computation/correctness_terminating.lean | 3 +++ src/analysis/convex/partition_of_unity.lean | 3 +++ src/analysis/normed/group/controlled_closure.lean | 3 +++ src/analysis/normed_space/compact_operator.lean | 3 +++ src/category_theory/abelian/pseudoelements.lean | 3 +++ src/category_theory/preadditive/hom_orthogonal.lean | 3 +++ src/category_theory/sites/cover_preserving.lean | 3 +++ src/computability/halting.lean | 3 +++ src/computability/partrec_code.lean | 3 +++ src/dynamics/ergodic/ergodic.lean | 3 +++ src/dynamics/ergodic/measure_preserving.lean | 3 +++ src/measure_theory/covering/vitali_family.lean | 3 +++ src/measure_theory/decomposition/unsigned_hahn.lean | 3 +++ src/measure_theory/group/arithmetic.lean | 3 +++ src/measure_theory/lattice.lean | 3 +++ src/measure_theory/measure/mutually_singular.lean | 3 +++ src/measure_theory/measure/open_pos.lean | 3 +++ src/probability/conditional_probability.lean | 3 +++ src/topology/algebra/equicontinuity.lean | 3 +++ src/topology/category/Profinite/as_limit.lean | 3 +++ src/topology/category/Top/limits/products.lean | 3 +++ src/topology/continuous_function/locally_constant.lean | 3 +++ 24 files changed, 72 insertions(+) diff --git a/src/algebra/category/Group/epi_mono.lean b/src/algebra/category/Group/epi_mono.lean index 4cf1e4732ee70..c43176e057940 100644 --- a/src/algebra/category/Group/epi_mono.lean +++ b/src/algebra/category/Group/epi_mono.lean @@ -8,6 +8,9 @@ import group_theory.quotient_group /-! # Monomorphisms and epimorphisms in `Group` + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. In this file, we prove monomorphisms in category of group are injective homomorphisms and epimorphisms are surjective homomorphisms. -/ diff --git a/src/algebra/category/Group/equivalence_Group_AddGroup.lean b/src/algebra/category/Group/equivalence_Group_AddGroup.lean index 00d5f4ec48617..036c5d2170992 100644 --- a/src/algebra/category/Group/equivalence_Group_AddGroup.lean +++ b/src/algebra/category/Group/equivalence_Group_AddGroup.lean @@ -9,6 +9,9 @@ import algebra.hom.equiv.type_tags /-! # Equivalence between `Group` and `AddGroup` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains two equivalences: * `Group_AddGroup_equivalence` : the equivalence between `Group` and `AddGroup` by sending `X : Group` to `additive X` and `Y : AddGroup` to `multiplicative Y`. diff --git a/src/algebra/continued_fractions/computation/correctness_terminating.lean b/src/algebra/continued_fractions/computation/correctness_terminating.lean index 7f976b1e1d255..d8d086b85e8c3 100644 --- a/src/algebra/continued_fractions/computation/correctness_terminating.lean +++ b/src/algebra/continued_fractions/computation/correctness_terminating.lean @@ -12,6 +12,9 @@ import tactic.field_simp /-! # Correctness of Terminating Continued Fraction Computations (`generalized_continued_fraction.of`) +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Summary We show the correctness of the diff --git a/src/analysis/convex/partition_of_unity.lean b/src/analysis/convex/partition_of_unity.lean index de3e06952c621..1518123f88636 100644 --- a/src/analysis/convex/partition_of_unity.lean +++ b/src/analysis/convex/partition_of_unity.lean @@ -9,6 +9,9 @@ import analysis.convex.combination /-! # Partition of unity and convex sets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove the following lemma, see `exists_continuous_forall_mem_convex_of_local`. Let `X` be a normal paracompact topological space (e.g., any extended metric space). Let `E` be a topological real vector space. Let `t : X → set E` be a family of convex sets. Suppose that for each diff --git a/src/analysis/normed/group/controlled_closure.lean b/src/analysis/normed/group/controlled_closure.lean index b24944f2dea5f..e43218e83efbb 100644 --- a/src/analysis/normed/group/controlled_closure.lean +++ b/src/analysis/normed/group/controlled_closure.lean @@ -8,6 +8,9 @@ import analysis.specific_limits.normed /-! # Extending a backward bound on a normed group homomorphism from a dense set +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Possible TODO (from the PR's review, https://github.com/leanprover-community/mathlib/pull/8498 ): "This feels a lot like the second step in the proof of the Banach open mapping theorem (`exists_preimage_norm_le`) ... wonder if it would be possible to refactor it using one of [the diff --git a/src/analysis/normed_space/compact_operator.lean b/src/analysis/normed_space/compact_operator.lean index 793e39757aa21..3561553dd0bce 100644 --- a/src/analysis/normed_space/compact_operator.lean +++ b/src/analysis/normed_space/compact_operator.lean @@ -9,6 +9,9 @@ import topology.algebra.module.strong_topology /-! # Compact operators +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define compact linear operators between two topological vector spaces (TVS). ## Main definitions diff --git a/src/category_theory/abelian/pseudoelements.lean b/src/category_theory/abelian/pseudoelements.lean index f0c1701b3fc9f..bdccd6f16fd76 100644 --- a/src/category_theory/abelian/pseudoelements.lean +++ b/src/category_theory/abelian/pseudoelements.lean @@ -10,6 +10,9 @@ import algebra.category.Module.epi_mono /-! # Pseudoelements in abelian categories +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A *pseudoelement* of an object `X` in an abelian category `C` is an equivalence class of arrows ending in `X`, where two arrows are considered equivalent if we can find two epimorphisms with a common domain making a commutative square with the two arrows. While the construction shows that diff --git a/src/category_theory/preadditive/hom_orthogonal.lean b/src/category_theory/preadditive/hom_orthogonal.lean index 7da288e5423d0..3e8d5b54b5aa6 100644 --- a/src/category_theory/preadditive/hom_orthogonal.lean +++ b/src/category_theory/preadditive/hom_orthogonal.lean @@ -10,6 +10,9 @@ import linear_algebra.matrix.invariant_basis_number /-! # Hom orthogonal families. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A family of objects in a category with zero morphisms is "hom orthogonal" if the only morphism between distinct objects is the zero morphism. diff --git a/src/category_theory/sites/cover_preserving.lean b/src/category_theory/sites/cover_preserving.lean index 2b3133f5f98ab..fbfe8d21710ba 100644 --- a/src/category_theory/sites/cover_preserving.lean +++ b/src/category_theory/sites/cover_preserving.lean @@ -10,6 +10,9 @@ import tactic.apply_fun /-! # Cover-preserving functors between sites. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define cover-preserving functors between sites as functors that push covering sieves to covering sieves. A cover-preserving and compatible-preserving functor `G : C ⥤ D` then pulls sheaves on `D` back to sheaves on `C` via `G.op ⋙ -`. diff --git a/src/computability/halting.lean b/src/computability/halting.lean index 81357e4f27011..72f98378201bb 100644 --- a/src/computability/halting.lean +++ b/src/computability/halting.lean @@ -8,6 +8,9 @@ import computability.partrec_code /-! # Computability theory and the halting problem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A universal partial recursive function, Rice's theorem, and the halting problem. ## References diff --git a/src/computability/partrec_code.lean b/src/computability/partrec_code.lean index 766b429a9f17b..bd552ff740022 100644 --- a/src/computability/partrec_code.lean +++ b/src/computability/partrec_code.lean @@ -8,6 +8,9 @@ import computability.partrec /-! # Gödel Numbering for Partial Recursive Functions. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines `nat.partrec.code`, an inductive datatype describing code for partial recursive functions on ℕ. It defines an encoding for these codes, and proves that the constructors are primitive recursive with respect to the encoding. diff --git a/src/dynamics/ergodic/ergodic.lean b/src/dynamics/ergodic/ergodic.lean index 99eb0e4bedfa6..a9a005f971042 100644 --- a/src/dynamics/ergodic/ergodic.lean +++ b/src/dynamics/ergodic/ergodic.lean @@ -8,6 +8,9 @@ import dynamics.ergodic.measure_preserving /-! # Ergodic maps and measures +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Let `f : α → α` be measure preserving with respect to a measure `μ`. We say `f` is ergodic with respect to `μ` (or `μ` is ergodic with respect to `f`) if the only measurable sets `s` such that `f⁻¹' s = s` are either almost empty or full. diff --git a/src/dynamics/ergodic/measure_preserving.lean b/src/dynamics/ergodic/measure_preserving.lean index 8a11e4d7efa6d..9942a6e3fcffd 100644 --- a/src/dynamics/ergodic/measure_preserving.lean +++ b/src/dynamics/ergodic/measure_preserving.lean @@ -8,6 +8,9 @@ import measure_theory.measure.ae_measurable /-! # Measure preserving maps +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We say that `f : α → β` is a measure preserving map w.r.t. measures `μ : measure α` and `ν : measure β` if `f` is measurable and `map f μ = ν`. In this file we define the predicate `measure_theory.measure_preserving` and prove its basic properties. diff --git a/src/measure_theory/covering/vitali_family.lean b/src/measure_theory/covering/vitali_family.lean index a3b7d7b46cf2d..2104ee59faf9f 100644 --- a/src/measure_theory/covering/vitali_family.lean +++ b/src/measure_theory/covering/vitali_family.lean @@ -8,6 +8,9 @@ import measure_theory.measure.measure_space /-! # Vitali families +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + On a metric space `X` with a measure `μ`, consider for each `x : X` a family of measurable sets with nonempty interiors, called `sets_at x`. This family is a Vitali family if it satisfies the following property: consider a (possibly non-measurable) set `s`, and for any `x` in `s` a diff --git a/src/measure_theory/decomposition/unsigned_hahn.lean b/src/measure_theory/decomposition/unsigned_hahn.lean index 612c2f03d4633..1f6150bc2c764 100644 --- a/src/measure_theory/decomposition/unsigned_hahn.lean +++ b/src/measure_theory/decomposition/unsigned_hahn.lean @@ -8,6 +8,9 @@ import measure_theory.measure.measure_space /-! # Unsigned Hahn decomposition theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves the unsigned version of the Hahn decomposition theorem. ## Main statements diff --git a/src/measure_theory/group/arithmetic.lean b/src/measure_theory/group/arithmetic.lean index 3f35f14bcb471..64c9f1e7bd8e6 100644 --- a/src/measure_theory/group/arithmetic.lean +++ b/src/measure_theory/group/arithmetic.lean @@ -8,6 +8,9 @@ import measure_theory.measure.ae_measurable /-! # Typeclasses for measurability of operations +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define classes `has_measurable_mul` etc and prove dot-style lemmas (`measurable.mul`, `ae_measurable.mul` etc). For binary operations we define two typeclasses: diff --git a/src/measure_theory/lattice.lean b/src/measure_theory/lattice.lean index 46576ba07dbdf..24a1453a0d1ae 100644 --- a/src/measure_theory/lattice.lean +++ b/src/measure_theory/lattice.lean @@ -9,6 +9,9 @@ import measure_theory.measure.ae_measurable /-! # Typeclasses for measurability of lattice operations +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define classes `has_measurable_sup` and `has_measurable_inf` and prove dot-style lemmas (`measurable.sup`, `ae_measurable.sup` etc). For binary operations we define two typeclasses: diff --git a/src/measure_theory/measure/mutually_singular.lean b/src/measure_theory/measure/mutually_singular.lean index 0146f97cf9d92..a0744f3019349 100644 --- a/src/measure_theory/measure/mutually_singular.lean +++ b/src/measure_theory/measure/mutually_singular.lean @@ -7,6 +7,9 @@ import measure_theory.measure.measure_space /-! # Mutually singular measures +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Two measures `μ`, `ν` are said to be mutually singular (`measure_theory.measure.mutually_singular`, localized notation `μ ⟂ₘ ν`) if there exists a measurable set `s` such that `μ s = 0` and `ν sᶜ = 0`. The measurability of `s` is an unnecessary assumption (see diff --git a/src/measure_theory/measure/open_pos.lean b/src/measure_theory/measure/open_pos.lean index a0ef83afa07e0..08f94286ff3ee 100644 --- a/src/measure_theory/measure/open_pos.lean +++ b/src/measure_theory/measure/open_pos.lean @@ -8,6 +8,9 @@ import measure_theory.measure.measure_space /-! # Measures positive on nonempty opens +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define a typeclass for measures that are positive on nonempty opens, see `measure_theory.measure.is_open_pos_measure`. Examples include (additive) Haar measures, as well as measures that have positive density with respect to a Haar measure. We also prove some basic facts diff --git a/src/probability/conditional_probability.lean b/src/probability/conditional_probability.lean index d790fe4a57208..6fc49d2152536 100644 --- a/src/probability/conditional_probability.lean +++ b/src/probability/conditional_probability.lean @@ -8,6 +8,9 @@ import measure_theory.measure.measure_space /-! # Conditional Probability +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines conditional probability and includes basic results relating to it. Given some measure `μ` defined on a measure space on some type `Ω` and some `s : set Ω`, diff --git a/src/topology/algebra/equicontinuity.lean b/src/topology/algebra/equicontinuity.lean index c7fdcc25e2f3c..8b68b827aa406 100644 --- a/src/topology/algebra/equicontinuity.lean +++ b/src/topology/algebra/equicontinuity.lean @@ -7,6 +7,9 @@ import topology.algebra.uniform_convergence /-! # Algebra-related equicontinuity criterions + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ open function diff --git a/src/topology/category/Profinite/as_limit.lean b/src/topology/category/Profinite/as_limit.lean index 71d6757804488..9796c77c7ab6c 100644 --- a/src/topology/category/Profinite/as_limit.lean +++ b/src/topology/category/Profinite/as_limit.lean @@ -9,6 +9,9 @@ import topology.discrete_quotient /-! # Profinite sets as limits of finite sets. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We show that any profinite set is isomorphic to the limit of its discrete (hence finite) quotients. diff --git a/src/topology/category/Top/limits/products.lean b/src/topology/category/Top/limits/products.lean index fa5c7931f9f69..cee1c228483cb 100644 --- a/src/topology/category/Top/limits/products.lean +++ b/src/topology/category/Top/limits/products.lean @@ -9,6 +9,9 @@ import topology.category.Top.limits.basic /-! # Products and coproducts in the category of topological spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + -/ open topological_space diff --git a/src/topology/continuous_function/locally_constant.lean b/src/topology/continuous_function/locally_constant.lean index 6c77ca5939f66..215538f00f006 100644 --- a/src/topology/continuous_function/locally_constant.lean +++ b/src/topology/continuous_function/locally_constant.lean @@ -10,6 +10,9 @@ import topology.continuous_function.algebra /-! # The algebra morphism from locally constant functions to continuous functions. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + -/ namespace locally_constant From 2f8347015b12b0864dfaf366ec4909eb70c78740 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Thu, 11 May 2023 07:22:31 +0000 Subject: [PATCH 0992/1291] chore(probability/independence): create independence folder, split old file into two parts (#18987) Split `probability/independence` into two files, to put Kolmogorov's zero-one law in its own file. The two new files are placed in an `independence` folder. Also fix a typo in a name in `probability/borel_cantelli`. --- src/probability/borel_cantelli.lean | 6 +- src/probability/conditional_expectation.lean | 2 +- .../basic.lean} | 172 ---------------- src/probability/independence/zero_one.lean | 184 ++++++++++++++++++ src/probability/integration.lean | 2 +- 5 files changed, 189 insertions(+), 177 deletions(-) rename src/probability/{independence.lean => independence/basic.lean} (85%) create mode 100644 src/probability/independence/zero_one.lean diff --git a/src/probability/borel_cantelli.lean b/src/probability/borel_cantelli.lean index 928753b73658e..466f4da94528b 100644 --- a/src/probability/borel_cantelli.lean +++ b/src/probability/borel_cantelli.lean @@ -5,7 +5,7 @@ Authors: Kexing Ying -/ import probability.martingale.borel_cantelli import probability.conditional_expectation -import probability.independence +import probability.independence.basic /-! @@ -46,7 +46,7 @@ begin exact indep_supr_of_disjoint (λ k, (hf k).measurable.comap_le) hfi (by simpa), end -lemma Indep_fun.condexp_natrual_ae_eq_of_lt +lemma Indep_fun.condexp_natural_ae_eq_of_lt [second_countable_topology β] [complete_space β] [normed_space ℝ β] (hf : ∀ i, strongly_measurable (f i)) (hfi : Indep_fun (λ i, mβ) f μ) (hij : i < j) : μ[f j | filtration.natural f hf i] =ᵐ[μ] λ ω, μ[f j] := @@ -59,7 +59,7 @@ lemma Indep_set.condexp_indicator_filtration_of_set_ae_eq μ[(s j).indicator (λ ω, 1 : Ω → ℝ) | filtration_of_set hsm i] =ᵐ[μ] λ ω, (μ (s j)).to_real := begin rw filtration.filtration_of_set_eq_natural hsm, - refine (Indep_fun.condexp_natrual_ae_eq_of_lt _ hs.Indep_fun_indicator hij).trans _, + refine (Indep_fun.condexp_natural_ae_eq_of_lt _ hs.Indep_fun_indicator hij).trans _, { simp only [integral_indicator_const _ (hsm _), algebra.id.smul_eq_mul, mul_one] }, { apply_instance } end diff --git a/src/probability/conditional_expectation.lean b/src/probability/conditional_expectation.lean index bc789e927ee6a..d77627101bee3 100644 --- a/src/probability/conditional_expectation.lean +++ b/src/probability/conditional_expectation.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kexing Ying -/ import probability.notation -import probability.independence +import probability.independence.basic import measure_theory.function.conditional_expectation.basic /-! diff --git a/src/probability/independence.lean b/src/probability/independence/basic.lean similarity index 85% rename from src/probability/independence.lean rename to src/probability/independence/basic.lean index 1b2268f2a8889..d7827c9312e8b 100644 --- a/src/probability/independence.lean +++ b/src/probability/independence/basic.lean @@ -28,9 +28,6 @@ import measure_theory.constructions.pi * `Indep_sets.Indep`: if π-systems are independent as sets of sets, then the measurable space structures they generate are independent. * `indep_sets.indep`: variant with two π-systems. -* `measure_zero_or_one_of_measurable_set_limsup_at_top`: Kolmogorov's 0-1 law. Any set which is - measurable with respect to the tail σ-algebra `limsup s at_top` of an independent sequence of - σ-algebras `s` has probability 0 or 1. ## Implementation notes @@ -897,173 +894,4 @@ end end indep_fun - -/-! ### Kolmogorov's 0-1 law - -Let `s : ι → measurable_space Ω` be an independent sequence of sub-σ-algebras. Then any set which -is measurable with respect to the tail σ-algebra `limsup s at_top` has probability 0 or 1. --/ - -section zero_one_law - -variables {m m0 : measurable_space Ω} {μ : measure Ω} - -lemma measure_eq_zero_or_one_or_top_of_indep_set_self {t : set Ω} (h_indep : indep_set t t μ) : - μ t = 0 ∨ μ t = 1 ∨ μ t = ∞ := -begin - specialize h_indep t t (measurable_set_generate_from (set.mem_singleton t)) - (measurable_set_generate_from (set.mem_singleton t)), - by_cases h0 : μ t = 0, - { exact or.inl h0, }, - by_cases h_top : μ t = ∞, - { exact or.inr (or.inr h_top), }, - rw [← one_mul (μ (t ∩ t)), set.inter_self, ennreal.mul_eq_mul_right h0 h_top] at h_indep, - exact or.inr (or.inl h_indep.symm), -end - -lemma measure_eq_zero_or_one_of_indep_set_self [is_finite_measure μ] {t : set Ω} - (h_indep : indep_set t t μ) : - μ t = 0 ∨ μ t = 1 := -begin - have h_0_1_top := measure_eq_zero_or_one_or_top_of_indep_set_self h_indep, - simpa [measure_ne_top μ] using h_0_1_top, -end - -variables [is_probability_measure μ] {s : ι → measurable_space Ω} - -open filter - -lemma indep_bsupr_compl (h_le : ∀ n, s n ≤ m0) (h_indep : Indep s μ) (t : set ι) : - indep (⨆ n ∈ t, s n) (⨆ n ∈ tᶜ, s n) μ := -indep_supr_of_disjoint h_le h_indep disjoint_compl_right - -section abstract -variables {α : Type*} {p : set ι → Prop} {f : filter ι} {ns : α → set ι} - -/-! We prove a version of Kolmogorov's 0-1 law for the σ-algebra `limsup s f` where `f` is a filter -for which we can define the following two functions: -* `p : set ι → Prop` such that for a set `t`, `p t → tᶜ ∈ f`, -* `ns : α → set ι` a directed sequence of sets which all verify `p` and such that - `⋃ a, ns a = set.univ`. - -For the example of `f = at_top`, we can take `p = bdd_above` and `ns : ι → set ι := λ i, set.Iic i`. --/ - -lemma indep_bsupr_limsup (h_le : ∀ n, s n ≤ m0) (h_indep : Indep s μ) - (hf : ∀ t, p t → tᶜ ∈ f) {t : set ι} (ht : p t) : - indep (⨆ n ∈ t, s n) (limsup s f) μ := -begin - refine indep_of_indep_of_le_right (indep_bsupr_compl h_le h_indep t) _, - refine Limsup_le_of_le (by is_bounded_default) _, - simp only [set.mem_compl_iff, eventually_map], - exact eventually_of_mem (hf t ht) le_supr₂, -end - -lemma indep_supr_directed_limsup (h_le : ∀ n, s n ≤ m0) (h_indep : Indep s μ) - (hf : ∀ t, p t → tᶜ ∈ f) (hns : directed (≤) ns) (hnsp : ∀ a, p (ns a)) : - indep (⨆ a, ⨆ n ∈ (ns a), s n) (limsup s f) μ := -begin - refine indep_supr_of_directed_le _ _ _ _, - { exact λ a, indep_bsupr_limsup h_le h_indep hf (hnsp a), }, - { exact λ a, supr₂_le (λ n hn, h_le n), }, - { exact limsup_le_supr.trans (supr_le h_le), }, - { intros a b, - obtain ⟨c, hc⟩ := hns a b, - refine ⟨c, _, _⟩; refine supr_mono (λ n, supr_mono' (λ hn, ⟨_, le_rfl⟩)), - { exact hc.1 hn, }, - { exact hc.2 hn, }, }, -end - -lemma indep_supr_limsup (h_le : ∀ n, s n ≤ m0) (h_indep : Indep s μ) (hf : ∀ t, p t → tᶜ ∈ f) - (hns : directed (≤) ns) (hnsp : ∀ a, p (ns a)) (hns_univ : ∀ n, ∃ a, n ∈ ns a) : - indep (⨆ n, s n) (limsup s f) μ := -begin - suffices : (⨆ a, ⨆ n ∈ (ns a), s n) = ⨆ n, s n, - { rw ← this, - exact indep_supr_directed_limsup h_le h_indep hf hns hnsp, }, - rw supr_comm, - refine supr_congr (λ n, _), - have : (⨆ (i : α) (H : n ∈ ns i), s n) = (⨆ (h : ∃ i, n ∈ ns i), s n), by rw supr_exists, - haveI : nonempty (∃ (i : α), n ∈ ns i) := ⟨hns_univ n⟩, - rw [this, supr_const], -end - -lemma indep_limsup_self (h_le : ∀ n, s n ≤ m0) (h_indep : Indep s μ) (hf : ∀ t, p t → tᶜ ∈ f) - (hns : directed (≤) ns) (hnsp : ∀ a, p (ns a)) (hns_univ : ∀ n, ∃ a, n ∈ ns a) : - indep (limsup s f) (limsup s f) μ := -indep_of_indep_of_le_left (indep_supr_limsup h_le h_indep hf hns hnsp hns_univ) limsup_le_supr - -theorem measure_zero_or_one_of_measurable_set_limsup (h_le : ∀ n, s n ≤ m0) (h_indep : Indep s μ) - (hf : ∀ t, p t → tᶜ ∈ f) (hns : directed (≤) ns) (hnsp : ∀ a, p (ns a)) - (hns_univ : ∀ n, ∃ a, n ∈ ns a) {t : set Ω} (ht_tail : measurable_set[limsup s f] t) : - μ t = 0 ∨ μ t = 1 := -measure_eq_zero_or_one_of_indep_set_self - ((indep_limsup_self h_le h_indep hf hns hnsp hns_univ).indep_set_of_measurable_set - ht_tail ht_tail) - -end abstract - -section at_top -variables [semilattice_sup ι] [no_max_order ι] [nonempty ι] - -lemma indep_limsup_at_top_self (h_le : ∀ n, s n ≤ m0) (h_indep : Indep s μ) : - indep (limsup s at_top) (limsup s at_top) μ := -begin - let ns : ι → set ι := set.Iic, - have hnsp : ∀ i, bdd_above (ns i) := λ i, bdd_above_Iic, - refine indep_limsup_self h_le h_indep _ _ hnsp _, - { simp only [mem_at_top_sets, ge_iff_le, set.mem_compl_iff, bdd_above, upper_bounds, - set.nonempty], - rintros t ⟨a, ha⟩, - obtain ⟨b, hb⟩ : ∃ b, a < b := exists_gt a, - refine ⟨b, λ c hc hct, _⟩, - suffices : ∀ i ∈ t, i < c, from lt_irrefl c (this c hct), - exact λ i hi, (ha hi).trans_lt (hb.trans_le hc), }, - { exact monotone.directed_le (λ i j hij k hki, le_trans hki hij), }, - { exact λ n, ⟨n, le_rfl⟩, }, -end - -/-- **Kolmogorov's 0-1 law** : any event in the tail σ-algebra of an independent sequence of -sub-σ-algebras has probability 0 or 1. -The tail σ-algebra `limsup s at_top` is the same as `⋂ n, ⋃ i ≥ n, s i`. -/ -theorem measure_zero_or_one_of_measurable_set_limsup_at_top (h_le : ∀ n, s n ≤ m0) - (h_indep : Indep s μ) {t : set Ω} (ht_tail : measurable_set[limsup s at_top] t) : - μ t = 0 ∨ μ t = 1 := -measure_eq_zero_or_one_of_indep_set_self - ((indep_limsup_at_top_self h_le h_indep).indep_set_of_measurable_set ht_tail ht_tail) - -end at_top - -section at_bot -variables [semilattice_inf ι] [no_min_order ι] [nonempty ι] - -lemma indep_limsup_at_bot_self (h_le : ∀ n, s n ≤ m0) (h_indep : Indep s μ) : - indep (limsup s at_bot) (limsup s at_bot) μ := -begin - let ns : ι → set ι := set.Ici, - have hnsp : ∀ i, bdd_below (ns i) := λ i, bdd_below_Ici, - refine indep_limsup_self h_le h_indep _ _ hnsp _, - { simp only [mem_at_bot_sets, ge_iff_le, set.mem_compl_iff, bdd_below, lower_bounds, - set.nonempty], - rintros t ⟨a, ha⟩, - obtain ⟨b, hb⟩ : ∃ b, b < a := exists_lt a, - refine ⟨b, λ c hc hct, _⟩, - suffices : ∀ i ∈ t, c < i, from lt_irrefl c (this c hct), - exact λ i hi, hc.trans_lt (hb.trans_le (ha hi)), }, - { exact directed_of_inf (λ i j hij k hki, hij.trans hki), }, - { exact λ n, ⟨n, le_rfl⟩, }, -end - -/-- **Kolmogorov's 0-1 law** : any event in the tail σ-algebra of an independent sequence of -sub-σ-algebras has probability 0 or 1. -/ -theorem measure_zero_or_one_of_measurable_set_limsup_at_bot (h_le : ∀ n, s n ≤ m0) - (h_indep : Indep s μ) {t : set Ω} (ht_tail : measurable_set[limsup s at_bot] t) : - μ t = 0 ∨ μ t = 1 := -measure_eq_zero_or_one_of_indep_set_self - ((indep_limsup_at_bot_self h_le h_indep).indep_set_of_measurable_set ht_tail ht_tail) - -end at_bot - -end zero_one_law - end probability_theory diff --git a/src/probability/independence/zero_one.lean b/src/probability/independence/zero_one.lean new file mode 100644 index 0000000000000..a806d92324928 --- /dev/null +++ b/src/probability/independence/zero_one.lean @@ -0,0 +1,184 @@ +/- +Copyright (c) 2021 Rémy Degenne. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Rémy Degenne +-/ +import probability.independence.basic + +/-! +# Kolmogorov's 0-1 law + +Let `s : ι → measurable_space Ω` be an independent sequence of sub-σ-algebras. Then any set which +is measurable with respect to the tail σ-algebra `limsup s at_top` has probability 0 or 1. + +## Main statements + +* `measure_zero_or_one_of_measurable_set_limsup_at_top`: Kolmogorov's 0-1 law. Any set which is + measurable with respect to the tail σ-algebra `limsup s at_top` of an independent sequence of + σ-algebras `s` has probability 0 or 1. +-/ + +open measure_theory measurable_space +open_locale measure_theory ennreal + +namespace probability_theory + +variables {Ω ι : Type*} {m m0 : measurable_space Ω} {μ : measure Ω} + +lemma measure_eq_zero_or_one_or_top_of_indep_set_self {t : set Ω} (h_indep : indep_set t t μ) : + μ t = 0 ∨ μ t = 1 ∨ μ t = ∞ := +begin + specialize h_indep t t (measurable_set_generate_from (set.mem_singleton t)) + (measurable_set_generate_from (set.mem_singleton t)), + by_cases h0 : μ t = 0, + { exact or.inl h0, }, + by_cases h_top : μ t = ∞, + { exact or.inr (or.inr h_top), }, + rw [← one_mul (μ (t ∩ t)), set.inter_self, ennreal.mul_eq_mul_right h0 h_top] at h_indep, + exact or.inr (or.inl h_indep.symm), +end + +lemma measure_eq_zero_or_one_of_indep_set_self [is_finite_measure μ] {t : set Ω} + (h_indep : indep_set t t μ) : + μ t = 0 ∨ μ t = 1 := +begin + have h_0_1_top := measure_eq_zero_or_one_or_top_of_indep_set_self h_indep, + simpa [measure_ne_top μ] using h_0_1_top, +end + +variables [is_probability_measure μ] {s : ι → measurable_space Ω} + +open filter + +lemma indep_bsupr_compl (h_le : ∀ n, s n ≤ m0) (h_indep : Indep s μ) (t : set ι) : + indep (⨆ n ∈ t, s n) (⨆ n ∈ tᶜ, s n) μ := +indep_supr_of_disjoint h_le h_indep disjoint_compl_right + +section abstract +variables {α : Type*} {p : set ι → Prop} {f : filter ι} {ns : α → set ι} + +/-! We prove a version of Kolmogorov's 0-1 law for the σ-algebra `limsup s f` where `f` is a filter +for which we can define the following two functions: +* `p : set ι → Prop` such that for a set `t`, `p t → tᶜ ∈ f`, +* `ns : α → set ι` a directed sequence of sets which all verify `p` and such that + `⋃ a, ns a = set.univ`. + +For the example of `f = at_top`, we can take `p = bdd_above` and `ns : ι → set ι := λ i, set.Iic i`. +-/ + +lemma indep_bsupr_limsup (h_le : ∀ n, s n ≤ m0) (h_indep : Indep s μ) + (hf : ∀ t, p t → tᶜ ∈ f) {t : set ι} (ht : p t) : + indep (⨆ n ∈ t, s n) (limsup s f) μ := +begin + refine indep_of_indep_of_le_right (indep_bsupr_compl h_le h_indep t) _, + refine Limsup_le_of_le (by is_bounded_default) _, + simp only [set.mem_compl_iff, eventually_map], + exact eventually_of_mem (hf t ht) le_supr₂, +end + +lemma indep_supr_directed_limsup (h_le : ∀ n, s n ≤ m0) (h_indep : Indep s μ) + (hf : ∀ t, p t → tᶜ ∈ f) (hns : directed (≤) ns) (hnsp : ∀ a, p (ns a)) : + indep (⨆ a, ⨆ n ∈ (ns a), s n) (limsup s f) μ := +begin + refine indep_supr_of_directed_le _ _ _ _, + { exact λ a, indep_bsupr_limsup h_le h_indep hf (hnsp a), }, + { exact λ a, supr₂_le (λ n hn, h_le n), }, + { exact limsup_le_supr.trans (supr_le h_le), }, + { intros a b, + obtain ⟨c, hc⟩ := hns a b, + refine ⟨c, _, _⟩; refine supr_mono (λ n, supr_mono' (λ hn, ⟨_, le_rfl⟩)), + { exact hc.1 hn, }, + { exact hc.2 hn, }, }, +end + +lemma indep_supr_limsup (h_le : ∀ n, s n ≤ m0) (h_indep : Indep s μ) (hf : ∀ t, p t → tᶜ ∈ f) + (hns : directed (≤) ns) (hnsp : ∀ a, p (ns a)) (hns_univ : ∀ n, ∃ a, n ∈ ns a) : + indep (⨆ n, s n) (limsup s f) μ := +begin + suffices : (⨆ a, ⨆ n ∈ (ns a), s n) = ⨆ n, s n, + { rw ← this, + exact indep_supr_directed_limsup h_le h_indep hf hns hnsp, }, + rw supr_comm, + refine supr_congr (λ n, _), + have : (⨆ (i : α) (H : n ∈ ns i), s n) = (⨆ (h : ∃ i, n ∈ ns i), s n), by rw supr_exists, + haveI : nonempty (∃ (i : α), n ∈ ns i) := ⟨hns_univ n⟩, + rw [this, supr_const], +end + +lemma indep_limsup_self (h_le : ∀ n, s n ≤ m0) (h_indep : Indep s μ) (hf : ∀ t, p t → tᶜ ∈ f) + (hns : directed (≤) ns) (hnsp : ∀ a, p (ns a)) (hns_univ : ∀ n, ∃ a, n ∈ ns a) : + indep (limsup s f) (limsup s f) μ := +indep_of_indep_of_le_left (indep_supr_limsup h_le h_indep hf hns hnsp hns_univ) limsup_le_supr + +theorem measure_zero_or_one_of_measurable_set_limsup (h_le : ∀ n, s n ≤ m0) (h_indep : Indep s μ) + (hf : ∀ t, p t → tᶜ ∈ f) (hns : directed (≤) ns) (hnsp : ∀ a, p (ns a)) + (hns_univ : ∀ n, ∃ a, n ∈ ns a) {t : set Ω} (ht_tail : measurable_set[limsup s f] t) : + μ t = 0 ∨ μ t = 1 := +measure_eq_zero_or_one_of_indep_set_self + ((indep_limsup_self h_le h_indep hf hns hnsp hns_univ).indep_set_of_measurable_set + ht_tail ht_tail) + +end abstract + +section at_top +variables [semilattice_sup ι] [no_max_order ι] [nonempty ι] + +lemma indep_limsup_at_top_self (h_le : ∀ n, s n ≤ m0) (h_indep : Indep s μ) : + indep (limsup s at_top) (limsup s at_top) μ := +begin + let ns : ι → set ι := set.Iic, + have hnsp : ∀ i, bdd_above (ns i) := λ i, bdd_above_Iic, + refine indep_limsup_self h_le h_indep _ _ hnsp _, + { simp only [mem_at_top_sets, ge_iff_le, set.mem_compl_iff, bdd_above, upper_bounds, + set.nonempty], + rintros t ⟨a, ha⟩, + obtain ⟨b, hb⟩ : ∃ b, a < b := exists_gt a, + refine ⟨b, λ c hc hct, _⟩, + suffices : ∀ i ∈ t, i < c, from lt_irrefl c (this c hct), + exact λ i hi, (ha hi).trans_lt (hb.trans_le hc), }, + { exact monotone.directed_le (λ i j hij k hki, le_trans hki hij), }, + { exact λ n, ⟨n, le_rfl⟩, }, +end + +/-- **Kolmogorov's 0-1 law** : any event in the tail σ-algebra of an independent sequence of +sub-σ-algebras has probability 0 or 1. +The tail σ-algebra `limsup s at_top` is the same as `⋂ n, ⋃ i ≥ n, s i`. -/ +theorem measure_zero_or_one_of_measurable_set_limsup_at_top (h_le : ∀ n, s n ≤ m0) + (h_indep : Indep s μ) {t : set Ω} (ht_tail : measurable_set[limsup s at_top] t) : + μ t = 0 ∨ μ t = 1 := +measure_eq_zero_or_one_of_indep_set_self + ((indep_limsup_at_top_self h_le h_indep).indep_set_of_measurable_set ht_tail ht_tail) + +end at_top + +section at_bot +variables [semilattice_inf ι] [no_min_order ι] [nonempty ι] + +lemma indep_limsup_at_bot_self (h_le : ∀ n, s n ≤ m0) (h_indep : Indep s μ) : + indep (limsup s at_bot) (limsup s at_bot) μ := +begin + let ns : ι → set ι := set.Ici, + have hnsp : ∀ i, bdd_below (ns i) := λ i, bdd_below_Ici, + refine indep_limsup_self h_le h_indep _ _ hnsp _, + { simp only [mem_at_bot_sets, ge_iff_le, set.mem_compl_iff, bdd_below, lower_bounds, + set.nonempty], + rintros t ⟨a, ha⟩, + obtain ⟨b, hb⟩ : ∃ b, b < a := exists_lt a, + refine ⟨b, λ c hc hct, _⟩, + suffices : ∀ i ∈ t, c < i, from lt_irrefl c (this c hct), + exact λ i hi, hc.trans_lt (hb.trans_le (ha hi)), }, + { exact directed_of_inf (λ i j hij k hki, hij.trans hki), }, + { exact λ n, ⟨n, le_rfl⟩, }, +end + +/-- **Kolmogorov's 0-1 law** : any event in the tail σ-algebra of an independent sequence of +sub-σ-algebras has probability 0 or 1. -/ +theorem measure_zero_or_one_of_measurable_set_limsup_at_bot (h_le : ∀ n, s n ≤ m0) + (h_indep : Indep s μ) {t : set Ω} (ht_tail : measurable_set[limsup s at_bot] t) : + μ t = 0 ∨ μ t = 1 := +measure_eq_zero_or_one_of_indep_set_self + ((indep_limsup_at_bot_self h_le h_indep).indep_set_of_measurable_set ht_tail ht_tail) + +end at_bot + +end probability_theory diff --git a/src/probability/integration.lean b/src/probability/integration.lean index d889e879b0b2e..78a17d30219de 100644 --- a/src/probability/integration.lean +++ b/src/probability/integration.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Martin Zinkevich, Vincent Beffara -/ import measure_theory.integral.set_integral -import probability.independence +import probability.independence.basic /-! # Integration in Probability Theory From 8c1b484d6a214e059531e22f1be9898ed6c1fd47 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 11 May 2023 21:30:27 +0000 Subject: [PATCH 0993/1291] feat(topology/sets/compacts): add `positive_compacts.map` (#18872) Also adds some missing functorial lemmas about `map`. --- src/topology/sets/compacts.lean | 60 ++++++++++++++++++++++++++++----- 1 file changed, 52 insertions(+), 8 deletions(-) diff --git a/src/topology/sets/compacts.lean b/src/topology/sets/compacts.lean index e4a4b96b9fb58..8dc271235107e 100644 --- a/src/topology/sets/compacts.lean +++ b/src/topology/sets/compacts.lean @@ -26,7 +26,7 @@ For a topological space `α`, open set -variables {α β : Type*} [topological_space α] [topological_space β] +variables {α β γ : Type*} [topological_space α] [topological_space β] [topological_space γ] namespace topological_space @@ -93,20 +93,35 @@ end protected def map (f : α → β) (hf : continuous f) (K : compacts α) : compacts β := ⟨f '' K.1, K.2.image hf⟩ -@[simp] lemma coe_map {f : α → β} (hf : continuous f) (s : compacts α) : +@[simp, norm_cast] lemma coe_map {f : α → β} (hf : continuous f) (s : compacts α) : (s.map f hf : set β) = f '' s := rfl +@[simp] lemma map_id (K : compacts α) : K.map id continuous_id = K := compacts.ext $ set.image_id _ + +lemma map_comp (f : β → γ) (g : α → β) (hf : continuous f) (hg : continuous g) (K : compacts α) : + K.map (f ∘ g) (hf.comp hg) = (K.map g hg).map f hf := compacts.ext $ set.image_comp _ _ _ + /-- A homeomorphism induces an equivalence on compact sets, by taking the image. -/ -@[simp] protected def equiv (f : α ≃ₜ β) : compacts α ≃ compacts β := +@[simps] protected def equiv (f : α ≃ₜ β) : compacts α ≃ compacts β := { to_fun := compacts.map f f.continuous, inv_fun := compacts.map _ f.symm.continuous, left_inv := λ s, by { ext1, simp only [coe_map, ← image_comp, f.symm_comp_self, image_id] }, right_inv := λ s, by { ext1, simp only [coe_map, ← image_comp, f.self_comp_symm, image_id] } } +@[simp] lemma equiv_refl : compacts.equiv (homeomorph.refl α) = equiv.refl _ := +equiv.ext map_id + +@[simp] lemma equiv_trans (f : α ≃ₜ β) (g : β ≃ₜ γ) : + compacts.equiv (f.trans g) = (compacts.equiv f).trans (compacts.equiv g) := +equiv.ext $ map_comp _ _ _ _ + +@[simp] lemma equiv_symm (f : α ≃ₜ β) : compacts.equiv f.symm = (compacts.equiv f).symm := +rfl + /-- The image of a compact set under a homeomorphism can also be expressed as a preimage. -/ -lemma equiv_to_fun_val (f : α ≃ₜ β) (K : compacts α) : - (compacts.equiv f K).1 = f.symm ⁻¹' K.1 := -congr_fun (image_eq_preimage_of_inverse f.left_inv f.right_inv) K.1 +lemma coe_equiv_apply_eq_preimage (f : α ≃ₜ β) (K : compacts α) : + (compacts.equiv f K : set β) = f.symm ⁻¹' (K : set α) := +f.to_equiv.image_eq_preimage K /-- The product of two `compacts`, as a `compacts` in the product space. -/ protected def prod (K : compacts α) (L : compacts β) : compacts (α × β) := @@ -225,6 +240,26 @@ order_top.lift (coe : _ → set α) (λ _ _, id) rfl @[simp] lemma coe_top [compact_space α] [nonempty α] : (↑(⊤ : positive_compacts α) : set α) = univ := rfl +/-- The image of a positive compact set under a continuous open map. -/ +protected def map (f : α → β) (hf : continuous f) (hf' : is_open_map f) (K : positive_compacts α) : + positive_compacts β := +{ interior_nonempty' := + (K.interior_nonempty'.image _).mono (hf'.image_interior_subset K.to_compacts), + ..K.map f hf } + +@[simp, norm_cast] lemma coe_map {f : α → β} (hf : continuous f) (hf' : is_open_map f) + (s : positive_compacts α) : + (s.map f hf hf' : set β) = f '' s := rfl + +@[simp] lemma map_id (K : positive_compacts α) : K.map id continuous_id is_open_map.id = K := +positive_compacts.ext $ set.image_id _ + +lemma map_comp (f : β → γ) (g : α → β) (hf : continuous f) (hg : continuous g) + (hf' : is_open_map f) (hg' : is_open_map g) + (K : positive_compacts α) : + K.map (f ∘ g) (hf.comp hg) (hf'.comp hg') = (K.map g hg hg').map f hf hf' := +positive_compacts.ext $ set.image_comp _ _ _ + lemma _root_.exists_positive_compacts_subset [locally_compact_space α] {U : set α} (ho : is_open U) (hn : U.nonempty) : ∃ K : positive_compacts α, ↑K ⊆ U := let ⟨x, hx⟩ := hn, ⟨K, hKc, hxK, hKU⟩ := exists_compact_subset ho hx in ⟨⟨⟨K, hKc⟩, ⟨x, hxK⟩⟩, hKU⟩ @@ -322,8 +357,17 @@ instance : inhabited (compact_opens α) := ⟨⊥⟩ compact_opens β := ⟨s.to_compacts.map f hf, hf' _ s.is_open⟩ -@[simp] lemma coe_map {f : α → β} (hf : continuous f) (hf' : is_open_map f) (s : compact_opens α) : - (s.map f hf hf' : set β) = f '' s := rfl +@[simp, norm_cast] lemma coe_map {f : α → β} (hf : continuous f) (hf' : is_open_map f) + (s : compact_opens α) : (s.map f hf hf' : set β) = f '' s := rfl + +@[simp] lemma map_id (K : compact_opens α) : K.map id continuous_id is_open_map.id = K := +compact_opens.ext $ set.image_id _ + +lemma map_comp (f : β → γ) (g : α → β) (hf : continuous f) (hg : continuous g) + (hf' : is_open_map f) (hg' : is_open_map g) + (K : compact_opens α) : + K.map (f ∘ g) (hf.comp hg) (hf'.comp hg') = (K.map g hg hg').map f hf hf' := +compact_opens.ext $ set.image_comp _ _ _ /-- The product of two `compact_opens`, as a `compact_opens` in the product space. -/ protected def prod (K : compact_opens α) (L : compact_opens β) : From b599f4e4e5cf1fbcb4194503671d3d9e569c1fce Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Fri, 12 May 2023 05:44:02 +0000 Subject: [PATCH 0994/1291] chore(algebra/order/group/defs): remove some instances that cause timeouts with etaExperiment (#18985) Seeing what CI thinks about removing these instances. This is a backport of https://github.com/leanprover-community/mathlib4/pull/3905 Co-authored-by: Scott Morrison --- src/algebra/order/group/defs.lean | 8 ++++++-- 1 file changed, 6 insertions(+), 2 deletions(-) diff --git a/src/algebra/order/group/defs.lean b/src/algebra/order/group/defs.lean index 69d9fe8ae12a7..cd13a15851640 100644 --- a/src/algebra/order/group/defs.lean +++ b/src/algebra/order/group/defs.lean @@ -55,15 +55,19 @@ instance ordered_comm_group.to_ordered_cancel_comm_monoid [ordered_comm_group α example (α : Type u) [ordered_add_comm_group α] : covariant_class α α (swap (+)) (<) := add_right_cancel_semigroup.covariant_swap_add_lt_of_covariant_swap_add_le α +-- Backporting note: this instance is not used, +-- and causes timeouts when interacting with etaExperiment. /-- A choice-free shortcut instance. -/ @[to_additive "A choice-free shortcut instance."] -instance ordered_comm_group.to_contravariant_class_left_le (α : Type u) [ordered_comm_group α] : +theorem ordered_comm_group.to_contravariant_class_left_le (α : Type u) [ordered_comm_group α] : contravariant_class α α (*) (≤) := { elim := λ a b c bc, by simpa using mul_le_mul_left' bc a⁻¹, } +-- Backporting note: this instance is not used, +-- and causes timeouts when interacting with etaExperiment. /-- A choice-free shortcut instance. -/ @[to_additive "A choice-free shortcut instance."] -instance ordered_comm_group.to_contravariant_class_right_le (α : Type u) [ordered_comm_group α] : +theorem ordered_comm_group.to_contravariant_class_right_le (α : Type u) [ordered_comm_group α] : contravariant_class α α (swap (*)) (≤) := { elim := λ a b c bc, by simpa using mul_le_mul_right' bc a⁻¹, } From 9e6d4aec88203d856eb35204b26d306896fd3399 Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Fri, 12 May 2023 08:46:25 +0000 Subject: [PATCH 0995/1291] feat(analysis/mellin_transform): Mellin transforms (#18822) This PR defines the Mellin transform of a locally integrable function, and proves it is holomorphic on a suitable vertical strip. --- src/analysis/mellin_transform.lean | 318 ++++++++++++++++++++++++ src/analysis/special_functions/pow.lean | 14 ++ 2 files changed, 332 insertions(+) create mode 100644 src/analysis/mellin_transform.lean diff --git a/src/analysis/mellin_transform.lean b/src/analysis/mellin_transform.lean new file mode 100644 index 0000000000000..fc32cd4d4c28f --- /dev/null +++ b/src/analysis/mellin_transform.lean @@ -0,0 +1,318 @@ +/- +Copyright (c) 2023 David Loeffler. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: David Loeffler +-/ + +import analysis.special_functions.improper_integrals +import analysis.calculus.parametric_integral + +/-! # The Mellin transform + +We define the Mellin transform of a locally integrable function on `Ioi 0`, and show it is +differentiable in a suitable vertical strip. + +## Main statements + +- `mellin` : the Mellin transform `∫ (t : ℝ) in Ioi 0, t ^ (s - 1) • f t`, + where `s` is a complex number. +- `mellin_differentiable_at_of_is_O_rpow` : if `f` is `O(x ^ (-a))` at infinity, and + `O(x ^ (-b))` at 0, then `mellin f` is holomorphic on the domain `b < re s < a`. + +-/ + +open measure_theory set filter asymptotics topological_space + +open_locale topology + +noncomputable theory + +section defs + +variables {E : Type*} [normed_add_comm_group E] + +/-- The Mellin transform of a function `f` (for a complex exponent `s`), defined as the integral of +`t ^ (s - 1) • f` over `Ioi 0`. -/ +def mellin [normed_space ℂ E] [complete_space E] (f : ℝ → E) (s : ℂ) : E := +∫ t : ℝ in Ioi 0, (t : ℂ) ^ (s - 1) • f t + +end defs + +open real complex (hiding exp log abs_of_nonneg) + +variables {E : Type*} [normed_add_comm_group E] + +section mellin_convergent +/-! ## Convergence of Mellin transform integrals -/ + +/-- Auxiliary lemma to reduce convergence statements from vector-valued functions to real +scalar-valued functions. -/ +lemma mellin_convergent_iff_norm [normed_space ℂ E] {f : ℝ → E} + {T : set ℝ} (hT : T ⊆ Ioi 0) (hT' : measurable_set T) + (hfc : ae_strongly_measurable f $ volume.restrict $ Ioi 0) {s : ℂ} : + integrable_on (λ t : ℝ, (t : ℂ) ^ (s - 1) • f t) T + ↔ integrable_on (λ t : ℝ, t ^ (s.re - 1) * ‖f t‖) T := +begin + have : ae_strongly_measurable (λ t : ℝ, (t : ℂ) ^ (s - 1) • f t) (volume.restrict T), + { refine ((continuous_at.continuous_on _).ae_strongly_measurable hT').smul (hfc.mono_set hT), + exact λ t ht, continuous_at_of_real_cpow_const _ _ (or.inr $ ne_of_gt (hT ht)) }, + rw [integrable_on, ←integrable_norm_iff this, ←integrable_on], + refine integrable_on_congr_fun (λ t ht, _) hT', + simp_rw [norm_smul, complex.norm_eq_abs, abs_cpow_eq_rpow_re_of_pos (hT ht), sub_re, one_re], +end + +/-- If `f` is a locally integrable real-valued function which is `O(x ^ (-a))` at `∞`, then for any +`s < a`, its Mellin transform converges on some neighbourhood of `+∞`. -/ +lemma mellin_convergent_top_of_is_O + {f : ℝ → ℝ} (hfc : ae_strongly_measurable f $ volume.restrict (Ioi 0)) + {a s : ℝ} (hf : is_O at_top f (λ t, t ^ (-a))) (hs : s < a) : + ∃ (c : ℝ), 0 < c ∧ integrable_on (λ t : ℝ, t ^ (s - 1) * f t) (Ioi c) := +begin + obtain ⟨d, hd, hd'⟩ := hf.exists_pos, + simp_rw [is_O_with, eventually_at_top] at hd', + obtain ⟨e, he⟩ := hd', + have he' : 0 < max e 1, from zero_lt_one.trans_le (le_max_right _ _), + refine ⟨max e 1, he', _, _⟩, + { refine ae_strongly_measurable.mul _ (hfc.mono_set (Ioi_subset_Ioi he'.le)), + refine (continuous_at.continuous_on (λ t ht, _)).ae_strongly_measurable measurable_set_Ioi, + exact continuous_at_rpow_const _ _ (or.inl $ (he'.trans ht).ne') }, + { have : ∀ᵐ (t : ℝ) ∂volume.restrict (Ioi $ max e 1), + ‖t ^ (s - 1) * f t‖ ≤ t ^ ((s - 1) + -a) * d, + { refine (ae_restrict_iff' measurable_set_Ioi).mpr (ae_of_all _ (λ t ht, _)), + have ht' : 0 < t, from he'.trans ht, + rw [norm_mul, rpow_add ht', ←norm_of_nonneg (rpow_nonneg_of_nonneg ht'.le (-a)), + mul_assoc, mul_comm _ d, norm_of_nonneg (rpow_nonneg_of_nonneg ht'.le _)], + exact mul_le_mul_of_nonneg_left (he t ((le_max_left e 1).trans_lt ht).le) + (rpow_pos_of_pos ht' _).le }, + refine (has_finite_integral.mul_const _ _).mono' this, + exact (integrable_on_Ioi_rpow_of_lt (by linarith) he').has_finite_integral } +end + +/-- If `f` is a locally integrable real-valued function which is `O(x ^ (-b))` at `0`, then for any +`b < s`, its Mellin transform converges on some right neighbourhood of `0`. -/ +lemma mellin_convergent_zero_of_is_O + {b : ℝ} {f : ℝ → ℝ} (hfc : ae_strongly_measurable f $ volume.restrict (Ioi 0)) + (hf : is_O (𝓝[Ioi 0] 0) f (λ t, t ^ (-b))) {s : ℝ} (hs : b < s) : + ∃ (c : ℝ), 0 < c ∧ integrable_on (λ t : ℝ, t ^ (s - 1) * f t) (Ioc 0 c) := +begin + obtain ⟨d, hd, hd'⟩ := hf.exists_pos, + simp_rw [is_O_with, eventually_nhds_within_iff, metric.eventually_nhds_iff, gt_iff_lt] at hd', + obtain ⟨ε, hε, hε'⟩ := hd', + refine ⟨ε, hε, integrable_on_Ioc_iff_integrable_on_Ioo.mpr ⟨_, _⟩⟩, + { refine ae_strongly_measurable.mul _ (hfc.mono_set Ioo_subset_Ioi_self), + refine (continuous_at.continuous_on (λ t ht, _)).ae_strongly_measurable measurable_set_Ioo, + exact continuous_at_rpow_const _ _ (or.inl ht.1.ne') }, + { apply has_finite_integral.mono', + { show has_finite_integral (λ t, d * t ^ (s - b - 1)) _, + refine (integrable.has_finite_integral _).const_mul _, + rw [←integrable_on, ←integrable_on_Ioc_iff_integrable_on_Ioo, + ←interval_integrable_iff_integrable_Ioc_of_le hε.le], + exact interval_integral.interval_integrable_rpow' (by linarith) }, + { refine (ae_restrict_iff' measurable_set_Ioo).mpr (eventually_of_forall $ λ t ht, _), + rw [mul_comm, norm_mul], + specialize hε' _ ht.1, + { rw [dist_eq_norm, sub_zero, norm_of_nonneg (le_of_lt ht.1)], + exact ht.2 }, + { refine (mul_le_mul_of_nonneg_right hε' (norm_nonneg _)).trans _, + simp_rw [norm_of_nonneg (rpow_nonneg_of_nonneg (le_of_lt ht.1) _), mul_assoc], + refine mul_le_mul_of_nonneg_left (le_of_eq _) hd.le, + rw ←rpow_add ht.1, + congr' 1, + abel } } }, +end + +/-- If `f` is a locally integrable real-valued function on `Ioi 0` which is `O(x ^ (-a))` at `∞` +and `O(x ^ (-b))` at `0`, then its Mellin transform integral converges for `b < s < a`. -/ +lemma mellin_convergent_of_is_O_scalar + {a b : ℝ} {f : ℝ → ℝ} {s : ℝ} + (hfc : locally_integrable_on f $ Ioi 0) + (hf_top : is_O at_top f (λ t, t ^ (-a))) (hs_top : s < a) + (hf_bot : is_O (𝓝[Ioi 0] 0) f (λ t, t ^ (-b))) (hs_bot : b < s) : + integrable_on (λ t : ℝ, t ^ (s - 1) * f t) (Ioi 0) := +begin + obtain ⟨c1, hc1, hc1'⟩ := mellin_convergent_top_of_is_O hfc.ae_strongly_measurable hf_top hs_top, + obtain ⟨c2, hc2, hc2'⟩ := mellin_convergent_zero_of_is_O hfc.ae_strongly_measurable hf_bot hs_bot, + have : Ioi 0 = Ioc 0 c2 ∪ Ioc c2 c1 ∪ Ioi c1, + { rw [union_assoc, Ioc_union_Ioi (le_max_right _ _), Ioc_union_Ioi + ((min_le_left _ _).trans (le_max_right _ _)), min_eq_left (lt_min hc2 hc1).le] }, + rw [this, integrable_on_union, integrable_on_union], + refine ⟨⟨hc2', integrable_on_Icc_iff_integrable_on_Ioc.mp _⟩, hc1'⟩, + refine (hfc.continuous_on_mul _ is_open_Ioi).integrable_on_compact_subset + (λ t ht, (hc2.trans_le ht.1 : 0 < t)) is_compact_Icc, + exact continuous_at.continuous_on (λ t ht, continuous_at_rpow_const _ _ $ or.inl $ ne_of_gt ht), +end + +lemma mellin_convergent_of_is_O_rpow [normed_space ℂ E] + {a b : ℝ} {f : ℝ → E} {s : ℂ} + (hfc : locally_integrable_on f $ Ioi 0) + (hf_top : is_O at_top f (λ t, t ^ (-a))) (hs_top : s.re < a) + (hf_bot : is_O (𝓝[Ioi 0] 0) f (λ t, t ^ (-b))) (hs_bot : b < s.re) : + integrable_on (λ t : ℝ, (t : ℂ) ^ (s - 1) • f t) (Ioi 0) := +begin + rw mellin_convergent_iff_norm (subset_refl _) measurable_set_Ioi + hfc.ae_strongly_measurable, + exact mellin_convergent_of_is_O_scalar + hfc.norm hf_top.norm_left hs_top hf_bot.norm_left hs_bot, +end + +end mellin_convergent + +section mellin_diff + +/-- If `f` is `O(x ^ (-a))` as `x → +∞`, then `log • f` is `O(x ^ (-b))` for every `b < a`. -/ +lemma is_O_rpow_top_log_smul [normed_space ℝ E] {a b : ℝ} {f : ℝ → E} + (hab : b < a) (hf : is_O at_top f (λ t, t ^ (-a))) : + is_O at_top (λ t : ℝ, log t • f t) (λ t, t ^ (-b)) := +begin + refine ((is_o_log_rpow_at_top (sub_pos.mpr hab)).is_O.smul hf).congr' + (eventually_of_forall (λ t, by refl)) + ((eventually_gt_at_top 0).mp (eventually_of_forall (λ t ht, _))), + rw [smul_eq_mul, ←rpow_add ht, ←sub_eq_add_neg, sub_eq_add_neg a, add_sub_cancel'], +end + +/-- If `f` is `O(x ^ (-a))` as `x → 0`, then `log • f` is `O(x ^ (-b))` for every `a < b`. -/ +lemma is_O_rpow_zero_log_smul [normed_space ℝ E] {a b : ℝ} {f : ℝ → E} + (hab : a < b) (hf : is_O (𝓝[Ioi 0] 0) f (λ t, t ^ (-a))) : + is_O (𝓝[Ioi 0] 0) (λ t : ℝ, log t • f t) (λ t, t ^ (-b)) := +begin + have : is_o (𝓝[Ioi 0] 0) log (λ t : ℝ, t ^ (a - b)), + { refine ((is_o_log_rpow_at_top (sub_pos.mpr hab)).neg_left.comp_tendsto + tendsto_inv_zero_at_top).congr' + (eventually_nhds_within_iff.mpr $ eventually_of_forall (λ t ht, _)) + (eventually_nhds_within_iff.mpr $ eventually_of_forall (λ t ht, _)), + { simp_rw [function.comp_app, ←one_div, log_div one_ne_zero (ne_of_gt ht), real.log_one, + zero_sub, neg_neg] }, + { simp_rw [function.comp_app, inv_rpow (le_of_lt ht), ←rpow_neg (le_of_lt ht), neg_sub] } }, + refine (this.is_O.smul hf).congr' + (eventually_of_forall (λ t, by refl)) + (eventually_nhds_within_iff.mpr (eventually_of_forall (λ t ht, _))), + simp_rw [smul_eq_mul, ←rpow_add ht], + congr' 1, + abel, +end + +/-- Suppose `f` is locally integrable on `(0, ∞)`, is `O(x ^ (-a))` as `x → ∞`, and is +`O(x ^ (-b))` as `x → 0`. Then its Mellin transform is differentiable on the domain `b < re s < a`, +with derivative equal to the Mellin transform of `log • f`. -/ +theorem mellin_has_deriv_of_is_O_rpow [complete_space E] [normed_space ℂ E] + {a b : ℝ} {f : ℝ → E} {s : ℂ} + (hfc : locally_integrable_on f $ Ioi 0) + (hf_top : is_O at_top f (λ t, t ^ (-a))) (hs_top : s.re < a) + (hf_bot : is_O (𝓝[Ioi 0] 0) f (λ t, t ^ (-b))) (hs_bot : b < s.re) : + has_deriv_at (mellin f) (mellin (λ t, (log t : ℂ) • f t) s) s := +begin + let F : ℂ → ℝ → E := λ z t, (t : ℂ) ^ (z - 1) • f t, + let F' : ℂ → ℝ → E := λ z t, ((t : ℂ) ^ (z - 1) * log t) • f t, + have hab : b < a := hs_bot.trans hs_top, + -- A convenient radius of ball within which we can uniformly bound the derivative. + obtain ⟨v, hv0, hv1, hv2⟩ : ∃ (v : ℝ), (0 < v) ∧ (v < s.re - b) ∧ (v < a - s.re), + { obtain ⟨w, hw1, hw2⟩ := exists_between (sub_pos.mpr hs_top), + obtain ⟨w', hw1', hw2'⟩ := exists_between (sub_pos.mpr hs_bot), + exact ⟨min w w', lt_min hw1 hw1', + (min_le_right _ _).trans_lt hw2', (min_le_left _ _).trans_lt hw2⟩ }, + let bound : ℝ → ℝ := λ t : ℝ, (t ^ (s.re + v - 1) + t ^ (s.re - v - 1)) * |log t| * ‖f t‖, + have h1 : ∀ᶠ (z : ℂ) in 𝓝 s, ae_strongly_measurable (F z) (volume.restrict $ Ioi 0), + { refine eventually_of_forall (λ z, ae_strongly_measurable.smul _ hfc.ae_strongly_measurable), + refine continuous_on.ae_strongly_measurable _ measurable_set_Ioi, + refine continuous_at.continuous_on (λ t ht, _), + exact (continuous_at_of_real_cpow_const _ _ (or.inr $ ne_of_gt ht)), }, + have h2 : integrable_on (F s) (Ioi 0), + { exact mellin_convergent_of_is_O_rpow hfc hf_top hs_top hf_bot hs_bot }, + have h3 : ae_strongly_measurable (F' s) (volume.restrict $ Ioi 0), + { apply locally_integrable_on.ae_strongly_measurable, + refine hfc.continuous_on_smul is_open_Ioi ((continuous_at.continuous_on (λ t ht, _)).mul _), + { exact continuous_at_of_real_cpow_const _ _ (or.inr $ ne_of_gt ht) }, + { refine continuous_of_real.comp_continuous_on _, + exact continuous_on_log.mono (subset_compl_singleton_iff.mpr not_mem_Ioi_self) } }, + have h4 : (∀ᵐ (t : ℝ) ∂volume.restrict (Ioi 0), ∀ (z : ℂ), + z ∈ metric.ball s v → ‖F' z t‖ ≤ bound t), + { refine (ae_restrict_iff' measurable_set_Ioi).mpr (ae_of_all _ $ λ t ht z hz, _), + simp_rw [bound, F', norm_smul, norm_mul, complex.norm_eq_abs (log _), complex.abs_of_real, + mul_assoc], + refine mul_le_mul_of_nonneg_right _ (mul_nonneg (abs_nonneg _) (norm_nonneg _)), + rw [complex.norm_eq_abs, abs_cpow_eq_rpow_re_of_pos ht], + rcases le_or_lt 1 t, + { refine le_add_of_le_of_nonneg (rpow_le_rpow_of_exponent_le h _) + (rpow_nonneg_of_nonneg (zero_le_one.trans h) _), + rw [sub_re, one_re, sub_le_sub_iff_right], + rw [mem_ball_iff_norm, complex.norm_eq_abs] at hz, + have hz' := (re_le_abs _).trans hz.le, + rwa [sub_re, sub_le_iff_le_add'] at hz' }, + { refine le_add_of_nonneg_of_le (rpow_pos_of_pos ht _).le + (rpow_le_rpow_of_exponent_ge ht h.le _), + rw [sub_re, one_re, sub_le_iff_le_add, sub_add_cancel], + rw [mem_ball_iff_norm', complex.norm_eq_abs] at hz, + have hz' := (re_le_abs _).trans hz.le, + rwa [sub_re, sub_le_iff_le_add, ←sub_le_iff_le_add'] at hz', } }, + have h5 : integrable_on bound (Ioi 0), + { simp_rw [bound, add_mul, mul_assoc], + suffices : ∀ {j : ℝ} (hj : b < j) (hj' : j < a), + integrable_on (λ (t : ℝ), t ^ (j - 1) * (|log t| * ‖f t‖)) (Ioi 0) volume, + { refine integrable.add (this _ _) (this _ _), + all_goals { linarith } }, + { intros j hj hj', + obtain ⟨w, hw1, hw2⟩ := exists_between hj, + obtain ⟨w', hw1', hw2'⟩ := exists_between hj', + refine mellin_convergent_of_is_O_scalar _ _ hw1' _ hw2, + { simp_rw mul_comm, + refine hfc.norm.mul_continuous_on _ is_open_Ioi, + refine continuous.comp_continuous_on continuous_abs (continuous_on_log.mono _), + exact subset_compl_singleton_iff.mpr not_mem_Ioi_self }, + { refine (is_O_rpow_top_log_smul hw2' hf_top).norm_left.congr' _ (eventually_eq.refl _ _), + refine (eventually_gt_at_top 0).mp (eventually_of_forall (λ t ht, _)), + simp only [norm_smul, real.norm_eq_abs] }, + { refine (is_O_rpow_zero_log_smul hw1 hf_bot).norm_left.congr' _ (eventually_eq.refl _ _), + refine eventually_nhds_within_iff.mpr (eventually_of_forall (λ t ht, _)), + simp only [norm_smul, real.norm_eq_abs] } } }, + have h6 : ∀ᵐ (t : ℝ) ∂volume.restrict (Ioi 0), ∀ (y : ℂ), + y ∈ metric.ball s v → has_deriv_at (λ (z : ℂ), F z t) (F' y t) y, + { dsimp only [F, F'], + refine (ae_restrict_iff' measurable_set_Ioi).mpr (ae_of_all _ $ λ t ht y hy, _), + have ht' : (t : ℂ) ≠ 0 := of_real_ne_zero.mpr (ne_of_gt ht), + have u1 : has_deriv_at (λ z : ℂ, (t : ℂ) ^ (z - 1)) (t ^ (y - 1) * log t) y, + { convert ((has_deriv_at_id' y).sub_const 1).const_cpow (or.inl ht') using 1, + rw of_real_log (le_of_lt ht), + ring }, + exact u1.smul_const (f t) }, + simpa only [F', mellin, mul_smul] using + (has_deriv_at_integral_of_dominated_loc_of_deriv_le hv0 h1 h2 h3 h4 h5 h6).2, +end + +/-- Suppose `f` is locally integrable on `(0, ∞)`, is `O(x ^ (-a))` as `x → ∞`, and is +`O(x ^ (-b))` as `x → 0`. Then its Mellin transform is differentiable on the domain `b < re s < a`. +-/ +lemma mellin_differentiable_at_of_is_O_rpow [complete_space E] [normed_space ℂ E] + {a b : ℝ} {f : ℝ → E} {s : ℂ} + (hfc : locally_integrable_on f $ Ioi 0) + (hf_top : is_O at_top f (λ t, t ^ (-a))) (hs_top : s.re < a) + (hf_bot : is_O (𝓝[Ioi 0] 0) f (λ t, t ^ (-b))) (hs_bot : b < s.re) : + differentiable_at ℂ (mellin f) s := +(mellin_has_deriv_of_is_O_rpow hfc hf_top hs_top hf_bot hs_bot).differentiable_at + +end mellin_diff + +section exp_decay + +/-- If `f` is locally integrable, decays exponentially at infinity, and is `O(x ^ (-b))` at 0, then +its Mellin transform converges for `b < s.re`. -/ +lemma mellin_convergent_of_is_O_rpow_exp [normed_space ℂ E] + {a b : ℝ} (ha : 0 < a) {f : ℝ → E} {s : ℂ} + (hfc : locally_integrable_on f $ Ioi 0) + (hf_top : is_O at_top f (λ t, exp (-a * t))) + (hf_bot : is_O (𝓝[Ioi 0] 0) f (λ t, t ^ (-b))) (hs_bot : b < s.re) : + integrable_on (λ t : ℝ, (t : ℂ) ^ (s - 1) • f t) (Ioi 0) := +mellin_convergent_of_is_O_rpow hfc (hf_top.trans (is_o_exp_neg_mul_rpow_at_top ha _).is_O) + (lt_add_one _) hf_bot hs_bot + +/-- If `f` is locally integrable, decays exponentially at infinity, and is `O(x ^ (-b))` at 0, then +its Mellin transform is holomorphic on `b < s.re`. -/ +lemma mellin_differentiable_at_of_is_O_rpow_exp [complete_space E] [normed_space ℂ E] + {a b : ℝ} (ha : 0 < a) {f : ℝ → E} {s : ℂ} + (hfc : locally_integrable_on f $ Ioi 0) + (hf_top : is_O at_top f (λ t, exp (-a * t))) + (hf_bot : is_O (𝓝[Ioi 0] 0) f (λ t, t ^ (-b))) (hs_bot : b < s.re) : + differentiable_at ℂ (mellin f) s := +mellin_differentiable_at_of_is_O_rpow hfc (hf_top.trans (is_o_exp_neg_mul_rpow_at_top ha _).is_O) + (lt_add_one _) hf_bot hs_bot + +end exp_decay diff --git a/src/analysis/special_functions/pow.lean b/src/analysis/special_functions/pow.lean index a91dee0dfda9f..21d391547585d 100644 --- a/src/analysis/special_functions/pow.lean +++ b/src/analysis/special_functions/pow.lean @@ -1180,6 +1180,20 @@ by simpa using is_o_zpow_exp_pos_mul_at_top k hb lemma is_o_rpow_exp_at_top (s : ℝ) : (λ x : ℝ, x ^ s) =o[at_top] exp := by simpa only [one_mul] using is_o_rpow_exp_pos_mul_at_top s one_pos +/-- `exp (-a * x) = o(x ^ s)` as `x → ∞`, for any positive `a` and real `s`. -/ +lemma is_o_exp_neg_mul_rpow_at_top {a : ℝ} (ha : 0 < a) (b : ℝ) : + is_o at_top (λ x : ℝ, exp (-a * x)) (λ x : ℝ, x ^ b) := +begin + apply is_o_of_tendsto', + { refine (eventually_gt_at_top 0).mp (eventually_of_forall $ λ t ht h, _), + rw rpow_eq_zero_iff_of_nonneg ht.le at h, + exact (ht.ne' h.1).elim }, + { refine (tendsto_exp_mul_div_rpow_at_top (-b) a ha).inv_tendsto_at_top.congr' _, + refine (eventually_ge_at_top 0).mp (eventually_of_forall $ λ t ht, _), + dsimp only, + rw [pi.inv_apply, inv_div, ←inv_div_inv, neg_mul, real.exp_neg, rpow_neg ht, inv_inv] } +end + lemma is_o_log_rpow_at_top {r : ℝ} (hr : 0 < r) : log =o[at_top] (λ x, x ^ r) := calc log =O[at_top] (λ x, r * log x) : is_O_self_const_mul _ hr.ne' _ _ ... =ᶠ[at_top] (λ x, log (x ^ r)) : From 066ecdb4834c7a4693e0f0e5154935a6f3d3f90c Mon Sep 17 00:00:00 2001 From: lukemantle Date: Fri, 12 May 2023 08:46:27 +0000 Subject: [PATCH 0996/1291] feat(ring_theory/polynomial): show gaussian form of `hermite n` (#18896) Show the equivalence of the polynomial form of the Hermite polynomials to the explicit form involving derivatives of the gaussian function. Co-authored-by: Jake Levinson Co-authored-by: lukemantle Co-authored-by: Eric Wieser Co-authored-by: Johan Commelin --- src/analysis/calculus/cont_diff.lean | 10 +++ .../{hermite.lean => hermite/basic.lean} | 2 + .../polynomial/hermite/gaussian.lean | 71 +++++++++++++++++++ 3 files changed, 83 insertions(+) rename src/ring_theory/polynomial/{hermite.lean => hermite/basic.lean} (96%) create mode 100644 src/ring_theory/polynomial/hermite/gaussian.lean diff --git a/src/analysis/calculus/cont_diff.lean b/src/analysis/calculus/cont_diff.lean index 45b377d3401ec..2a8cced40e013 100644 --- a/src/analysis/calculus/cont_diff.lean +++ b/src/analysis/calculus/cont_diff.lean @@ -2296,6 +2296,16 @@ lemma cont_diff.continuous_deriv (h : cont_diff 𝕜 n f₂) (hn : 1 ≤ n) : continuous (deriv f₂) := (cont_diff_succ_iff_deriv.mp (h.of_le hn)).2.continuous +lemma cont_diff.iterate_deriv : + ∀ (n : ℕ) {f₂ : 𝕜 → F} (hf : cont_diff 𝕜 ∞ f₂), cont_diff 𝕜 ∞ (deriv^[n] f₂) +| 0 f₂ hf := hf +| (n + 1) f₂ hf := cont_diff.iterate_deriv n (cont_diff_top_iff_deriv.mp hf).2 + +lemma cont_diff.iterate_deriv' (n : ℕ) : + ∀ (k : ℕ) {f₂ : 𝕜 → F} (hf : cont_diff 𝕜 (n + k : ℕ) f₂), cont_diff 𝕜 n (deriv^[k] f₂) +| 0 f₂ hf := hf +| (n + 1) f₂ hf := cont_diff.iterate_deriv' n (cont_diff_succ_iff_deriv.mp hf).2 + end deriv section restrict_scalars diff --git a/src/ring_theory/polynomial/hermite.lean b/src/ring_theory/polynomial/hermite/basic.lean similarity index 96% rename from src/ring_theory/polynomial/hermite.lean rename to src/ring_theory/polynomial/hermite/basic.lean index 92699c9b6594f..4a6d951db090c 100644 --- a/src/ring_theory/polynomial/hermite.lean +++ b/src/ring_theory/polynomial/hermite/basic.lean @@ -19,6 +19,7 @@ This file defines `polynomial.hermite n`, the nth probabilist's Hermite polynomi ## Results +* `polynomial.hermite_succ`: the recursion `hermite (n+1) = (x - d/dx) (hermite n)` * `polynomial.coeff_hermite_of_odd_add`: for `n`,`k` where `n+k` is odd, `(hermite n).coeff k` is zero. * `polynomial.monic_hermite`: for all `n`, `hermite n` is monic. @@ -39,6 +40,7 @@ noncomputable def hermite : ℕ → polynomial ℤ | 0 := 1 | (n+1) := X * (hermite n) - (hermite n).derivative +/-- The recursion `hermite (n+1) = (x - d/dx) (hermite n)` -/ @[simp] lemma hermite_succ (n : ℕ) : hermite (n+1) = X * (hermite n) - (hermite n).derivative := by rw hermite diff --git a/src/ring_theory/polynomial/hermite/gaussian.lean b/src/ring_theory/polynomial/hermite/gaussian.lean new file mode 100644 index 0000000000000..0dc26ea76c07f --- /dev/null +++ b/src/ring_theory/polynomial/hermite/gaussian.lean @@ -0,0 +1,71 @@ +/- +Copyright (c) 2023 Luke Mantle. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Luke Mantle, Jake Levinson +-/ + +import ring_theory.polynomial.hermite.basic +import analysis.special_functions.exp +import analysis.special_functions.exp_deriv + +/-! +# Hermite polynomials and Gaussians + +This file shows that the Hermite polynomial `hermite n` is (up to sign) the +polynomial factor occurring in the `n`th derivative of a gaussian. + +## Results + +* `polynomial.deriv_gaussian_eq_hermite_mul_gaussian`: + The Hermite polynomial is (up to sign) the polynomial factor occurring in the + `n`th derivative of a gaussian. + +## References + +* [Hermite Polynomials](https://en.wikipedia.org/wiki/Hermite_polynomials) + +-/ + +noncomputable theory +open polynomial + +namespace polynomial + +/-- `hermite n` is (up to sign) the factor appearing in `deriv^[n]` of a gaussian -/ +lemma deriv_gaussian_eq_hermite_mul_gaussian (n : ℕ) (x : ℝ) : + deriv^[n] (λ y, real.exp (-(y^2 / 2))) x = + (-1 : ℝ)^n * aeval x (hermite n) * real.exp (-(x^2 / 2)) := +begin + rw mul_assoc, + induction n with n ih generalizing x, + { rw [function.iterate_zero_apply, pow_zero, one_mul, hermite_zero, C_1, map_one, one_mul] }, + { replace ih : (deriv^[n] _) = _ := _root_.funext ih, + have deriv_gaussian : deriv (λ y, real.exp (-(y^2 / 2))) x = (-x) * real.exp (-(x^2 / 2)), + { simp [mul_comm, ← neg_mul] }, + rw [function.iterate_succ_apply', ih, deriv_const_mul_field, deriv_mul, pow_succ (-1 : ℝ), + deriv_gaussian, hermite_succ, map_sub, map_mul, aeval_X, polynomial.deriv_aeval], + ring, + { apply polynomial.differentiable_aeval }, + { simp } } +end + +lemma hermite_eq_deriv_gaussian (n : ℕ) (x : ℝ) : + aeval x (hermite n) = + (-1 : ℝ)^n * (deriv^[n] (λ y, real.exp (-(y^2 / 2))) x) / real.exp (-(x^2 / 2)) := +begin + rw deriv_gaussian_eq_hermite_mul_gaussian, + field_simp [real.exp_ne_zero], + rw [← @smul_eq_mul ℝ _ ((-1)^n), ← inv_smul_eq_iff₀, mul_assoc, smul_eq_mul, + ← inv_pow, ← neg_inv, inv_one], + exact pow_ne_zero _ (by norm_num), +end + +lemma hermite_eq_deriv_gaussian' (n : ℕ) (x : ℝ) : + aeval x (hermite n) = + (-1 : ℝ)^n * (deriv^[n] (λ y, real.exp (-(y^2 / 2))) x) * real.exp (x^2 / 2) := +begin + rw [hermite_eq_deriv_gaussian, real.exp_neg], + field_simp [real.exp_ne_zero], +end + +end polynomial From 343e80208d29d2d15f8050b929aa50fe4ce71b55 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 12 May 2023 08:46:28 +0000 Subject: [PATCH 0997/1291] chore(measure_theory/measure/outer_measure): lemmas about zero and top (#18983) --- src/measure_theory/measure/hausdorff.lean | 14 +++++++++++++ src/measure_theory/measure/measure_space.lean | 11 ++++++++++ src/measure_theory/measure/outer_measure.lean | 20 +++++++++++++++++++ 3 files changed, 45 insertions(+) diff --git a/src/measure_theory/measure/hausdorff.lean b/src/measure_theory/measure/hausdorff.lean index 2b8222aecfe00..2b28738a336d3 100644 --- a/src/measure_theory/measure/hausdorff.lean +++ b/src/measure_theory/measure/hausdorff.lean @@ -358,6 +358,14 @@ begin { simp [h0] } end +@[simp] lemma mk_metric_top : (mk_metric (λ _, ∞ : ℝ≥0∞ → ℝ≥0∞) : outer_measure X) = ⊤ := +begin + simp_rw [mk_metric, mk_metric', mk_metric'.pre, extend_top, bounded_by_top, eq_top_iff], + rw le_supr_iff, + intros b hb, + simpa using hb ⊤, +end + /-- If `m₁ d ≤ m₂ d` for `d < ε` for some `ε > 0` (we use `≤ᶠ[𝓝[≥] 0]` to state this), then `mk_metric m₁ hm₁ ≤ mk_metric m₂ hm₂`-/ lemma mk_metric_mono {m₁ m₂ : ℝ≥0∞ → ℝ≥0∞} (hle : m₁ ≤ᶠ[𝓝[≥] 0] m₂) : @@ -465,6 +473,12 @@ begin exact outer_measure.mk_metric_mono_smul hc h0 hle s end +@[simp] lemma mk_metric_top : (mk_metric (λ _, ∞ : ℝ≥0∞ → ℝ≥0∞) : measure X) = ⊤ := +begin + apply to_outer_measure_injective, + rw [mk_metric_to_outer_measure, outer_measure.mk_metric_top, to_outer_measure_top], +end + /-- If `m₁ d ≤ m₂ d` for `d < ε` for some `ε > 0` (we use `≤ᶠ[𝓝[≥] 0]` to state this), then `mk_metric m₁ hm₁ ≤ mk_metric m₂ hm₂`-/ lemma mk_metric_mono {m₁ m₂ : ℝ≥0∞ → ℝ≥0∞} (hle : m₁ ≤ᶠ[𝓝[≥] 0] m₂) : diff --git a/src/measure_theory/measure/measure_space.lean b/src/measure_theory/measure/measure_space.lean index c19686b5b94cb..9089d791ce8e2 100644 --- a/src/measure_theory/measure/measure_space.lean +++ b/src/measure_theory/measure/measure_space.lean @@ -938,6 +938,17 @@ instance [measurable_space α] : complete_lattice (measure α) := end Inf +@[simp] lemma _root_.measure_theory.outer_measure.to_measure_top [measurable_space α] : + (⊤ : outer_measure α).to_measure (by rw [outer_measure.top_caratheodory]; exact le_top) + = (⊤ : measure α) := +top_unique $ λ s hs, + by cases s.eq_empty_or_nonempty with h h; + simp [h, to_measure_apply ⊤ _ hs, outer_measure.top_apply] + +@[simp] lemma to_outer_measure_top [measurable_space α] : + (⊤ : measure α).to_outer_measure = (⊤ : outer_measure α) := +by rw [←outer_measure.to_measure_top, to_measure_to_outer_measure, outer_measure.trim_top] + @[simp] lemma top_add : ⊤ + μ = ⊤ := top_unique $ measure.le_add_right le_rfl @[simp] lemma add_top : μ + ⊤ = ⊤ := top_unique $ measure.le_add_left le_rfl diff --git a/src/measure_theory/measure/outer_measure.lean b/src/measure_theory/measure/outer_measure.lean index af24e0abc4afe..d4f4103e49130 100644 --- a/src/measure_theory/measure/outer_measure.lean +++ b/src/measure_theory/measure/outer_measure.lean @@ -704,6 +704,20 @@ theorem le_bounded_by' {μ : outer_measure α} : μ ≤ bounded_by m ↔ ∀ s : set α, s.nonempty → μ s ≤ m s := by { rw [le_bounded_by, forall_congr], intro s, cases s.eq_empty_or_nonempty with h h; simp [h] } +@[simp] lemma bounded_by_top : bounded_by (⊤ : set α → ℝ≥0∞) = ⊤ := +begin + rw [eq_top_iff, le_bounded_by'], + intros s hs, + rw top_apply hs, + exact le_rfl, +end + +@[simp] lemma bounded_by_zero : bounded_by (0 : set α → ℝ≥0∞) = 0 := +begin + rw [←coe_bot, eq_bot_iff], + apply bounded_by_le, +end + lemma smul_bounded_by {c : ℝ≥0∞} (hc : c ≠ ∞) : c • bounded_by m = bounded_by (c • m) := begin simp only [bounded_by, smul_of_function hc], @@ -1094,6 +1108,10 @@ lemma extend_congr {β : Type*} {Pb : β → Prop} {mb : Π s : β, Pb s → ℝ extend m sa = extend mb sb := infi_congr_Prop hP (λ h, hm _ _) +@[simp] lemma extend_top {α : Type*} {P : α → Prop} : + extend (λ s h, ∞ : Π (s : α), P s → ℝ≥0∞) = ⊤ := +funext $ λ x, infi_eq_top.mpr $ λ i, rfl + end extend section extend_set @@ -1337,6 +1355,8 @@ by simp [infi_subtype, infi_and, trim_eq_infi] theorem trim_trim (m : outer_measure α) : m.trim.trim = m.trim := trim_eq_trim_iff.2 $ λ s, m.trim_eq +@[simp] theorem trim_top : (⊤ : outer_measure α).trim = ⊤ := eq_top_iff.2 $ le_trim _ + @[simp] theorem trim_zero : (0 : outer_measure α).trim = 0 := ext $ λ s, le_antisymm (le_trans ((trim 0).mono (subset_univ s)) $ From bd15ff41b70f5e2cc210f26f25a8d5c53b20d3de Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Fri, 12 May 2023 08:46:29 +0000 Subject: [PATCH 0998/1291] chore(*): add mathlib4 synchronization comments (#18992) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.module.injective` * `category_theory.preadditive.yoneda.injective` * `category_theory.preadditive.yoneda.projective` * `category_theory.sites.cover_lifting` * `category_theory.sites.pushforward` * `computability.reduce` * `computability.tm_to_partrec` * `measure_theory.group.measurable_equiv` * `measure_theory.group.pointwise` * `probability.cond_count` * `probability.probability_mass_function.monad` * `topology.continuous_function.polynomial` --- src/algebra/module/injective.lean | 3 +++ src/category_theory/preadditive/yoneda/injective.lean | 3 +++ src/category_theory/preadditive/yoneda/projective.lean | 3 +++ src/category_theory/sites/cover_lifting.lean | 3 +++ src/category_theory/sites/pushforward.lean | 3 +++ src/computability/reduce.lean | 3 +++ src/computability/tm_to_partrec.lean | 3 +++ src/measure_theory/group/measurable_equiv.lean | 3 +++ src/measure_theory/group/pointwise.lean | 3 +++ src/probability/cond_count.lean | 3 +++ src/probability/probability_mass_function/monad.lean | 3 +++ src/topology/continuous_function/polynomial.lean | 3 +++ 12 files changed, 36 insertions(+) diff --git a/src/algebra/module/injective.lean b/src/algebra/module/injective.lean index 43b956e32afd4..cfbb97da496fc 100644 --- a/src/algebra/module/injective.lean +++ b/src/algebra/module/injective.lean @@ -12,6 +12,9 @@ import linear_algebra.linear_pmap /-! # Injective modules +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions * `module.injective`: an `R`-module `Q` is injective if and only if every injective `R`-linear diff --git a/src/category_theory/preadditive/yoneda/injective.lean b/src/category_theory/preadditive/yoneda/injective.lean index d26a0c76459f4..cd5a64ba95c78 100644 --- a/src/category_theory/preadditive/yoneda/injective.lean +++ b/src/category_theory/preadditive/yoneda/injective.lean @@ -9,6 +9,9 @@ import algebra.category.Group.epi_mono import algebra.category.Module.epi_mono /-! +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + An object is injective iff the preadditive yoneda functor on it preserves epimorphisms. -/ diff --git a/src/category_theory/preadditive/yoneda/projective.lean b/src/category_theory/preadditive/yoneda/projective.lean index 886bb6d2a5a87..0cc596a1ce183 100644 --- a/src/category_theory/preadditive/yoneda/projective.lean +++ b/src/category_theory/preadditive/yoneda/projective.lean @@ -9,6 +9,9 @@ import algebra.category.Group.epi_mono import algebra.category.Module.epi_mono /-! +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + An object is projective iff the preadditive coyoneda functor on it preserves epimorphisms. -/ diff --git a/src/category_theory/sites/cover_lifting.lean b/src/category_theory/sites/cover_lifting.lean index deca91fbd087b..e395e67985e3d 100644 --- a/src/category_theory/sites/cover_lifting.lean +++ b/src/category_theory/sites/cover_lifting.lean @@ -10,6 +10,9 @@ import category_theory.sites.cover_preserving /-! # Cover-lifting functors between sites. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define cover-lifting functors between sites as functors that pull covering sieves back to covering sieves. This concept is also known as *cocontinuous functors* or *cover-reflecting functors*, but we have chosen this name following [MM92] in order to avoid diff --git a/src/category_theory/sites/pushforward.lean b/src/category_theory/sites/pushforward.lean index 4d2545bd60937..c14c8316b1f0f 100644 --- a/src/category_theory/sites/pushforward.lean +++ b/src/category_theory/sites/pushforward.lean @@ -9,6 +9,9 @@ import category_theory.sites.left_exact /-! # Pushforward of sheaves +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions * `category_theory.sites.pushforward`: the induced functor `Sheaf J A ⥤ Sheaf K A` for a diff --git a/src/computability/reduce.lean b/src/computability/reduce.lean index 9eba50837e9f7..513b16f0280b1 100644 --- a/src/computability/reduce.lean +++ b/src/computability/reduce.lean @@ -8,6 +8,9 @@ import computability.halting /-! # Strong reducibility and degrees. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the notions of computable many-one reduction and one-one reduction between sets, and shows that the corresponding degrees form a semilattice. diff --git a/src/computability/tm_to_partrec.lean b/src/computability/tm_to_partrec.lean index 52bea45d745ee..267f6b2edda2b 100644 --- a/src/computability/tm_to_partrec.lean +++ b/src/computability/tm_to_partrec.lean @@ -11,6 +11,9 @@ import tactic.derive_fintype /-! # Modelling partial recursive functions using Turing machines +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines a simplified basis for partial recursive functions, and a `turing.TM2` model Turing machine for evaluating these functions. This amounts to a constructive proof that every `partrec` function can be evaluated by a Turing machine. diff --git a/src/measure_theory/group/measurable_equiv.lean b/src/measure_theory/group/measurable_equiv.lean index 4fda5629527a0..baa7f9a11cbb4 100644 --- a/src/measure_theory/group/measurable_equiv.lean +++ b/src/measure_theory/group/measurable_equiv.lean @@ -8,6 +8,9 @@ import measure_theory.group.arithmetic /-! # (Scalar) multiplication and (vector) addition as measurable equivalences +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define the following measurable equivalences: * `measurable_equiv.smul`: if a group `G` acts on `α` by measurable maps, then each element `c : G` diff --git a/src/measure_theory/group/pointwise.lean b/src/measure_theory/group/pointwise.lean index 86bbc5f745e23..82e55d4269208 100644 --- a/src/measure_theory/group/pointwise.lean +++ b/src/measure_theory/group/pointwise.lean @@ -8,6 +8,9 @@ import measure_theory.group.arithmetic /-! # Pointwise set operations on `measurable_set`s +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove several versions of the following fact: if `s` is a measurable set, then so is `a • s`. Note that the pointwise product of two measurable sets need not be measurable, so there is no `measurable_set.mul` etc. diff --git a/src/probability/cond_count.lean b/src/probability/cond_count.lean index 09a8ab8c12782..3145180686d02 100644 --- a/src/probability/cond_count.lean +++ b/src/probability/cond_count.lean @@ -8,6 +8,9 @@ import probability.conditional_probability /-! # Classical probability +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The classical formulation of probability states that the probability of an event occurring in a finite probability space is the ratio of that event to all possible events. This notion can be expressed with measure theory using diff --git a/src/probability/probability_mass_function/monad.lean b/src/probability/probability_mass_function/monad.lean index 08d2f85af463e..2110791e1bd61 100644 --- a/src/probability/probability_mass_function/monad.lean +++ b/src/probability/probability_mass_function/monad.lean @@ -8,6 +8,9 @@ import probability.probability_mass_function.basic /-! # Monad Operations for Probability Mass Functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file constructs two operations on `pmf` that give it a monad structure. `pure a` is the distribution where a single value `a` has probability `1`. `bind pa pb : pmf β` is the distribution given by sampling `a : α` from `pa : pmf α`, diff --git a/src/topology/continuous_function/polynomial.lean b/src/topology/continuous_function/polynomial.lean index e9707b52e3e12..a9c7142e3f834 100644 --- a/src/topology/continuous_function/polynomial.lean +++ b/src/topology/continuous_function/polynomial.lean @@ -10,6 +10,9 @@ import topology.unit_interval /-! # Constructions relating polynomial functions and continuous functions. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions * `polynomial.to_continuous_map_on p X`: for `X : set R`, interprets a polynomial `p` From ee7b9f9a9ac2a8d9f04ea39bbfe6b1a3be053b38 Mon Sep 17 00:00:00 2001 From: "Filippo A. E. Nuccio" Date: Fri, 12 May 2023 16:32:23 +0000 Subject: [PATCH 0999/1291] doc(algebra/euclidean_domain/defs): correct typos in doc (#19001) Correct two typos in the doc about the Main Statements about Euclidean Domains. The corresponding `mathlib-4` PR is [#3945. ](https://github.com/leanprover-community/mathlib4/pull/3945) --- src/algebra/euclidean_domain/defs.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/algebra/euclidean_domain/defs.lean b/src/algebra/euclidean_domain/defs.lean index 5d82d8bc4c87b..18d9d394c8437 100644 --- a/src/algebra/euclidean_domain/defs.lean +++ b/src/algebra/euclidean_domain/defs.lean @@ -32,10 +32,10 @@ don't satisfy the classical notion were provided independently by Hiblot and Nag ## Main statements -See `algebra.euclidean_domain.basic` for most of the theorems about Eucliean domains, +See `algebra.euclidean_domain.basic` for most of the theorems about Euclidean domains, including Bézout's lemma. -See `algebra.euclidean_domain.instances` for that facts that `ℤ` is a Euclidean domain, +See `algebra.euclidean_domain.instances` for the fact that `ℤ` is a Euclidean domain, as is any field. ## Notation From 403190b5419b3f03f1a2893ad9352ca7f7d8bff6 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Fri, 12 May 2023 18:49:43 +0000 Subject: [PATCH 1000/1291] feat(algebra/category/Module/change_of_rings): extension and restriction of scalars are adjoint (#15564) --- .../category/Module/change_of_rings.lean | 172 ++++++++++++++++++ 1 file changed, 172 insertions(+) diff --git a/src/algebra/category/Module/change_of_rings.lean b/src/algebra/category/Module/change_of_rings.lean index 9d8c285ba5110..fcf032e3c20f0 100644 --- a/src/algebra/category/Module/change_of_rings.lean +++ b/src/algebra/category/Module/change_of_rings.lean @@ -20,6 +20,11 @@ import ring_theory.tensor_product module structure is defined by `s • (s' ⊗ m) := (s * s') ⊗ m` and `R`-linear map `l : M ⟶ M'` is sent to `S`-linear map `s ⊗ m ↦ s ⊗ l m : S ⨂ M ⟶ S ⨂ M'`. +## Main results + +* `category_theory.Module.extend_restrict_scalars_adj`: given commutative rings `R, S` and a ring + homomorphism `f : R →+* S`, the extension and restriction of scalars by `f` are adjoint functors. + ## List of notations Let `R, S` be rings and `f : R →+* S` * if `M` is an `R`-module, `s : S` and `m : M`, then `s ⊗ₜ[R, f] m` is the pure tensor @@ -64,6 +69,10 @@ def restrict_scalars {R : Type u₁} {S : Type u₂} [ring R] [ring S] (f : R map_id' := λ _, linear_map.ext $ λ m, rfl, map_comp' := λ _ _ _ g h, linear_map.ext $ λ m, rfl } +instance {R : Type u₁} {S : Type u₂} [comm_ring R] [comm_ring S] (f : R →+* S) : + category_theory.faithful (restrict_scalars.{v} f) := +{ map_injective' := λ _ _ _ _ h, linear_map.ext $ λ x, by simpa only using fun_like.congr_fun h x } + @[simp] lemma restrict_scalars.map_apply {R : Type u₁} {S : Type u₂} [ring R] [ring S] (f : R →+* S) {M M' : Module.{v} S} (g : M ⟶ M') (x) : (restrict_scalars f).map g x = g x := rfl @@ -161,4 +170,167 @@ variables {R : Type u₁} {S : Type u₂} [comm_ring R] [comm_ring S] (f : R → end extend_scalars +namespace extend_restrict_scalars_adj + +open_locale change_of_rings +open tensor_product + +variables {R : Type u₁} {S : Type u₂} [comm_ring R] [comm_ring S] (f : R →+* S) + +/-- +Given `R`-module X and `S`-module Y and a map `g : (extend_scalars f).obj X ⟶ Y`, i.e. `S`-linear +map `S ⨂ X → Y`, there is a `X ⟶ (restrict_scalars f).obj Y`, i.e. `R`-linear map `X ⟶ Y` by +`x ↦ g (1 ⊗ x)`. +-/ +@[simps] def hom_equiv.to_restrict_scalars {X Y} (g : (extend_scalars f).obj X ⟶ Y) : + X ⟶ (restrict_scalars f).obj Y := +{ to_fun := λ x, g $ (1 : S) ⊗ₜ[R, f] x, + map_add' := λ _ _, by rw [tmul_add, map_add], + map_smul' := λ r x, + begin + letI : module R S := module.comp_hom S f, + letI : module R Y := module.comp_hom Y f, + rw [ring_hom.id_apply, restrict_scalars.smul_def, ←linear_map.map_smul, tmul_smul], + congr, + end } + +/-- +Given `R`-module X and `S`-module Y and a map `X ⟶ (restrict_scalars f).obj Y`, i.e `R`-linear map +`X ⟶ Y`, there is a map `(extend_scalars f).obj X ⟶ Y`, i.e `S`-linear map `S ⨂ X → Y` by +`s ⊗ x ↦ s • g x`. +-/ +@[simps] def hom_equiv.from_extend_scalars {X Y} (g : X ⟶ (restrict_scalars f).obj Y) : + (extend_scalars f).obj X ⟶ Y := +begin + letI m1 : module R S := module.comp_hom S f, letI m2 : module R Y := module.comp_hom Y f, + refine ⟨λ z, tensor_product.lift ⟨λ s, ⟨_, _, _⟩, _, _⟩ z, _, _⟩, + { exact λ x, s • g x }, + { intros, rw [map_add, smul_add], }, + { intros, rw [ring_hom.id_apply, smul_comm, ←linear_map.map_smul], }, + { intros, ext, simp only [linear_map.coe_mk, linear_map.add_apply], rw ←add_smul, }, + { intros, ext, + simp only [linear_map.coe_mk, ring_hom.id_apply, linear_map.smul_apply, + restrict_scalars.smul_def, smul_eq_mul], + convert mul_smul _ _ _, }, + { intros, rw [map_add], }, + { intros r z, + rw [ring_hom.id_apply], + induction z using tensor_product.induction_on with x y x y ih1 ih2, + { simp only [smul_zero, map_zero], }, + { simp only [linear_map.coe_mk, extend_scalars.smul_tmul, lift.tmul, ←mul_smul], }, + { rw [smul_add, map_add, ih1, ih2, map_add, smul_add], }, }, +end + +/-- +Given `R`-module X and `S`-module Y, `S`-linear linear maps `(extend_scalars f).obj X ⟶ Y` +bijectively correspond to `R`-linear maps `X ⟶ (restrict_scalars f).obj Y`. +-/ +@[simps] +def hom_equiv {X Y} : ((extend_scalars f).obj X ⟶ Y) ≃ (X ⟶ (restrict_scalars f).obj Y) := +{ to_fun := hom_equiv.to_restrict_scalars f, + inv_fun := hom_equiv.from_extend_scalars f, + left_inv := λ g, begin + ext z, + induction z using tensor_product.induction_on with x s z1 z2 ih1 ih2, + { simp only [map_zero], }, + { erw tensor_product.lift.tmul, + simp only [linear_map.coe_mk], + change S at x, + erw [←linear_map.map_smul, extend_scalars.smul_tmul, mul_one x], }, + { rw [map_add, map_add, ih1, ih2], } + end, + right_inv := λ g, + begin + ext, + rw [hom_equiv.to_restrict_scalars_apply, hom_equiv.from_extend_scalars_apply, lift.tmul, + linear_map.coe_mk, linear_map.coe_mk], + convert one_smul _ _, + end } + +/-- +For any `R`-module X, there is a natural `R`-linear map from `X` to `X ⨂ S` by sending `x ↦ x ⊗ 1` +-/ +@[simps] def unit.map {X} : X ⟶ (extend_scalars f ⋙ restrict_scalars f).obj X := +{ to_fun := λ x, (1 : S) ⊗ₜ[R, f] x, + map_add' := λ x x', by { rw tensor_product.tmul_add, }, + map_smul' := λ r x, by { letI m1 : module R S := module.comp_hom S f, tidy } } + +/-- +The natural transformation from identity functor on `R`-module to the composition of extension and +restriction of scalars. +-/ +@[simps] def unit : 𝟭 (Module R) ⟶ extend_scalars f ⋙ restrict_scalars f := +{ app := λ _, unit.map f, naturality' := λ X X' g, by tidy } + +/-- +For any `S`-module Y, there is a natural `R`-linear map from `S ⨂ Y` to `Y` by +`s ⊗ y ↦ s • y` +-/ +@[simps] def counit.map {Y} : (restrict_scalars f ⋙ extend_scalars f).obj Y ⟶ Y := +begin + letI m1 : module R S := module.comp_hom S f, + letI m2 : module R Y := module.comp_hom Y f, + refine ⟨tensor_product.lift ⟨λ (s : S), ⟨λ (y : Y), s • y, smul_add _, _⟩, _, _⟩, _, _⟩, + { intros, rw [ring_hom.id_apply, restrict_scalars.smul_def, ←mul_smul, mul_comm, mul_smul, + restrict_scalars.smul_def], }, + { intros, ext, simp only [linear_map.add_apply, linear_map.coe_mk, add_smul], }, + { intros, ext, + simpa only [ring_hom.id_apply, linear_map.smul_apply, linear_map.coe_mk, + @restrict_scalars.smul_def _ _ _ _ f ⟨S⟩, smul_eq_mul, mul_smul], }, + { intros, rw [map_add], }, + { intros s z, + rw [ring_hom.id_apply], + induction z using tensor_product.induction_on with x s' z1 z2 ih1 ih2, + { simp only [smul_zero, map_zero], }, + { simp only [extend_scalars.smul_tmul, linear_map.coe_mk, tensor_product.lift.tmul, mul_smul] }, + { rw [smul_add, map_add, map_add, ih1, ih2, smul_add], } }, +end + +/-- +The natural transformation from the composition of restriction and extension of scalars to the +identity functor on `S`-module. +-/ +@[simps] def counit : (restrict_scalars f ⋙ extend_scalars f) ⟶ (𝟭 (Module S)) := +{ app := λ _, counit.map f, + naturality' := λ Y Y' g, + begin + ext z, induction z using tensor_product.induction_on, + { simp only [map_zero] }, + { simp only [category_theory.functor.comp_map, Module.coe_comp, function.comp_app, + extend_scalars.map_tmul, restrict_scalars.map_apply, counit.map_apply, lift.tmul, + linear_map.coe_mk, category_theory.functor.id_map, linear_map.map_smulₛₗ, + ring_hom.id_apply] }, + { simp only [map_add, *] }, + end } + +end extend_restrict_scalars_adj + +/-- +Given commutative rings `R, S` and a ring hom `f : R →+* S`, the extension and restriction of +scalars by `f` are adjoint to each other. +-/ +@[simps] +def extend_restrict_scalars_adj {R : Type u₁} {S : Type u₂} [comm_ring R] [comm_ring S] + (f : R →+* S) : extend_scalars f ⊣ restrict_scalars f := +{ hom_equiv := λ _ _, extend_restrict_scalars_adj.hom_equiv f, + unit := extend_restrict_scalars_adj.unit f, + counit := extend_restrict_scalars_adj.counit f, + hom_equiv_unit' := λ X Y g, linear_map.ext $ λ x, by simp, + hom_equiv_counit' := λ X Y g, linear_map.ext $ λ x, + begin + induction x using tensor_product.induction_on, + { simp only [map_zero]}, + { simp only [extend_restrict_scalars_adj.hom_equiv_symm_apply, linear_map.coe_mk, + extend_restrict_scalars_adj.hom_equiv.from_extend_scalars_apply, tensor_product.lift.tmul, + extend_restrict_scalars_adj.counit_app, Module.coe_comp, function.comp_app, + extend_scalars.map_tmul, extend_restrict_scalars_adj.counit.map_apply] }, + { simp only [map_add, *], } + end } + +instance {R : Type u₁} {S : Type u₂} [comm_ring R] [comm_ring S] (f : R →+* S) : + category_theory.is_left_adjoint (extend_scalars f) := ⟨_, extend_restrict_scalars_adj f⟩ + +instance {R : Type u₁} {S : Type u₂} [comm_ring R] [comm_ring S] (f : R →+* S) : + category_theory.is_right_adjoint (restrict_scalars f) := ⟨_, extend_restrict_scalars_adj f⟩ + end category_theory.Module From 58537299e922f9c77df76cb613910914a479c1f7 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Sat, 13 May 2023 10:36:52 +0000 Subject: [PATCH 1001/1291] chore(analysis/calculus/fderiv): move large file to dedicated subfolder, prepare for split (#19007) Co-authored-by: Eric Wieser --- src/analysis/box_integral/divergence_theorem.lean | 2 +- src/analysis/calculus/conformal/normed_space.lean | 2 +- src/analysis/calculus/cont_diff_def.lean | 2 +- src/analysis/calculus/deriv.lean | 2 +- src/analysis/calculus/{fderiv.lean => fderiv/basic.lean} | 0 5 files changed, 4 insertions(+), 4 deletions(-) rename src/analysis/calculus/{fderiv.lean => fderiv/basic.lean} (100%) diff --git a/src/analysis/box_integral/divergence_theorem.lean b/src/analysis/box_integral/divergence_theorem.lean index 97600b5182095..ea65ee8a4ed60 100644 --- a/src/analysis/box_integral/divergence_theorem.lean +++ b/src/analysis/box_integral/divergence_theorem.lean @@ -5,7 +5,7 @@ Authors: Yury Kudryashov -/ import analysis.box_integral.basic import analysis.box_integral.partition.additive -import analysis.calculus.fderiv +import analysis.calculus.fderiv.basic /-! # Divergence integral for Henstock-Kurzweil integral diff --git a/src/analysis/calculus/conformal/normed_space.lean b/src/analysis/calculus/conformal/normed_space.lean index 03ad42c84f295..a187e8103b8b7 100644 --- a/src/analysis/calculus/conformal/normed_space.lean +++ b/src/analysis/calculus/conformal/normed_space.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yourong Zang -/ import analysis.normed_space.conformal_linear_map -import analysis.calculus.fderiv +import analysis.calculus.fderiv.basic /-! # Conformal Maps diff --git a/src/analysis/calculus/cont_diff_def.lean b/src/analysis/calculus/cont_diff_def.lean index b65c429a34abf..51293f418f5bc 100644 --- a/src/analysis/calculus/cont_diff_def.lean +++ b/src/analysis/calculus/cont_diff_def.lean @@ -3,7 +3,7 @@ Copyright (c) 2019 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ -import analysis.calculus.fderiv +import analysis.calculus.fderiv.basic import analysis.calculus.formal_multilinear_series /-! diff --git a/src/analysis/calculus/deriv.lean b/src/analysis/calculus/deriv.lean index 20327a2f4cc7b..4b673fbaf198b 100644 --- a/src/analysis/calculus/deriv.lean +++ b/src/analysis/calculus/deriv.lean @@ -3,7 +3,7 @@ Copyright (c) 2019 Gabriel Ebner. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Gabriel Ebner, Sébastien Gouëzel -/ -import analysis.calculus.fderiv +import analysis.calculus.fderiv.basic import data.polynomial.algebra_map import data.polynomial.derivative import linear_algebra.affine_space.slope diff --git a/src/analysis/calculus/fderiv.lean b/src/analysis/calculus/fderiv/basic.lean similarity index 100% rename from src/analysis/calculus/fderiv.lean rename to src/analysis/calculus/fderiv/basic.lean From 1d650c2e131f500f3c17f33b4d19d2ea15987f2c Mon Sep 17 00:00:00 2001 From: Kevin Buzzard Date: Sat, 13 May 2023 12:04:10 +0000 Subject: [PATCH 1002/1291] chore(category_theory/sites/dense_subsite): minor refactor (#19002) Lean 4 pointed out that `include H` came slightly too early in `category_theory/sites_dense_subsite`; it doesn't need to be in the definition of https://leanprover-community.github.io/mathlib_docs/category_theory/sites/dense_subsite.html#category_theory.cover_dense.types.pushforward_family , for example. I removed both the `include`s in the file, which changes the types of at least one declaration (a superfluous `H` is no longer there). This file is not a leaf file -- it's imported by all the algebraic geometry hierarchy for example -- but mathlib still compiles with these edits and I propose making this change in the ported version of this file. --- src/category_theory/sites/dense_subsite.lean | 19 +++++++++++++++---- 1 file changed, 15 insertions(+), 4 deletions(-) diff --git a/src/category_theory/sites/dense_subsite.lean b/src/category_theory/sites/dense_subsite.lean index ceac0e5660ff7..d17ad4ce184ca 100644 --- a/src/category_theory/sites/dense_subsite.lean +++ b/src/category_theory/sites/dense_subsite.lean @@ -141,7 +141,6 @@ lemma sheaf_eq_amalgamation (ℱ : Sheaf K A) {X : A} {U : D} {T : sieve U} (hT) t = (ℱ.cond X T hT).amalgamate x hx := (ℱ.cond X T hT).is_separated_for x t _ h ((ℱ.cond X T hT).is_amalgamation hx) -include H variable [full G] namespace types variables {ℱ : Dᵒᵖ ⥤ Type v} {ℱ' : SheafOfTypes.{v} K} (α : G.op ⋙ ℱ ⟶ G.op ⋙ ℱ'.val) @@ -154,9 +153,11 @@ def pushforward_family {X} (x : ℱ.obj (op X)) : family_of_elements ℱ'.val (cover_by_image G X) := λ Y f hf, ℱ'.val.map hf.some.lift.op $ α.app (op _) (ℱ.map hf.some.map.op x : _) +include H + /-- (Implementation). The `pushforward_family` defined is compatible. -/ lemma pushforward_family_compatible {X} (x : ℱ.obj (op X)) : - (pushforward_family H α x).compatible := + (pushforward_family α x).compatible := begin intros Y₁ Y₂ Z g₁ g₂ f₁ f₂ h₁ h₂ e, apply H.ext, @@ -175,15 +176,17 @@ begin simp [e] end +omit H + /-- (Implementation). The morphism `ℱ(X) ⟶ ℱ'(X)` given by gluing the `pushforward_family`. -/ noncomputable def app_hom (X : D) : ℱ.obj (op X) ⟶ ℱ'.val.obj (op X) := λ x, (ℱ'.cond _ (H.is_cover X)).amalgamate - (pushforward_family H α x) + (pushforward_family α x) (pushforward_family_compatible H α x) @[simp] lemma pushforward_family_apply {X} (x : ℱ.obj (op X)) {Y : C} (f : G.obj Y ⟶ X) : - pushforward_family H α x f (presieve.in_cover_by_image G f) = α.app (op Y) (ℱ.map f.op x) := + pushforward_family α x f (presieve.in_cover_by_image G f) = α.app (op Y) (ℱ.map f.op x) := begin unfold pushforward_family, refine congr_fun _ x, @@ -283,6 +286,7 @@ def sheaf_coyoneda_hom (α : G.op ⋙ ℱ ⟶ G.op ⋙ ℱ'.val) : simp end } +include H /-- (Implementation). `sheaf_coyoneda_hom` but the order of the arguments of the functor are swapped. -/ @@ -300,6 +304,8 @@ begin exact congr_fun ((α.app X).naturality i) x }, end +omit H + /-- Given an natural transformation `G ⋙ ℱ ⟶ G ⋙ ℱ'` between presheaves of arbitrary category, where `G` is full and cover-dense, and `ℱ'` is a sheaf, we may obtain a natural transformation @@ -312,6 +318,7 @@ let α' := sheaf_yoneda_hom H α in { app := λ X, yoneda.preimage (α'.app X), naturality' := λ X Y f, yoneda.map_injective (by simpa using α'.naturality f) } +include H /-- Given an natural isomorphism `G ⋙ ℱ ≅ G ⋙ ℱ'` between presheaves of arbitrary category, where `G` is full and cover-dense, and `ℱ', ℱ` are sheaves, @@ -335,6 +342,7 @@ begin apply as_iso (sheaf_hom H i.hom), end +omit H /-- Given an natural isomorphism `G ⋙ ℱ ≅ G ⋙ ℱ'` between presheaves of arbitrary category, where `G` is full and cover-dense, and `ℱ', ℱ` are sheaves, @@ -404,6 +412,7 @@ def restrict_hom_equiv_hom : (G.op ⋙ ℱ ⟶ G.op ⋙ ℱ'.val) ≃ (ℱ ⟶ left_inv := sheaf_hom_restrict_eq H, right_inv := sheaf_hom_eq H } +include H /-- Given a full and cover-dense functor `G` and a natural transformation of sheaves `α : ℱ ⟶ ℱ'`, if the pullback of `α` along `G` is iso, then `α` is also iso. @@ -431,6 +440,8 @@ begin simp [eq] end +omit H + noncomputable instance sites.pullback.full [faithful G] (Hp : cover_preserving J K G) : full (sites.pullback A H.compatible_preserving Hp) := From 6cf5900728239efa287df7761ec2a1ac9cf39b29 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Sat, 13 May 2023 13:20:17 +0000 Subject: [PATCH 1003/1291] chore(*): add mathlib4 synchronization comments (#19005) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.homology.short_exact.abelian` * `algebra.modeq` * `algebra.ring.boolean_ring` * `analysis.locally_convex.continuous_of_bounded` * `analysis.normed_space.units` * `category_theory.abelian.diagram_lemmas.four` * `category_theory.abelian.generator` * `category_theory.idempotents.biproducts` * `data.is_R_or_C.basic` * `model_theory.definability` * `model_theory.order` * `model_theory.semantics` * `model_theory.substructures` * `topology.category.Top.limits.pullbacks` --- src/algebra/homology/short_exact/abelian.lean | 3 +++ src/algebra/modeq.lean | 3 +++ src/algebra/ring/boolean_ring.lean | 3 +++ src/analysis/locally_convex/continuous_of_bounded.lean | 3 +++ src/analysis/normed_space/units.lean | 3 +++ src/category_theory/abelian/diagram_lemmas/four.lean | 3 +++ src/category_theory/abelian/generator.lean | 3 +++ src/category_theory/idempotents/biproducts.lean | 3 +++ src/data/is_R_or_C/basic.lean | 3 +++ src/model_theory/definability.lean | 3 +++ src/model_theory/order.lean | 3 +++ src/model_theory/semantics.lean | 3 +++ src/model_theory/substructures.lean | 3 +++ src/topology/category/Top/limits/pullbacks.lean | 3 +++ 14 files changed, 42 insertions(+) diff --git a/src/algebra/homology/short_exact/abelian.lean b/src/algebra/homology/short_exact/abelian.lean index 47f08a30936c6..639dea8217be2 100644 --- a/src/algebra/homology/short_exact/abelian.lean +++ b/src/algebra/homology/short_exact/abelian.lean @@ -9,6 +9,9 @@ import category_theory.abelian.diagram_lemmas.four /-! # Short exact sequences in abelian categories +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In an abelian category, a left-split or right-split short exact sequence admits a splitting. -/ diff --git a/src/algebra/modeq.lean b/src/algebra/modeq.lean index 26f397630e57b..68291700694b0 100644 --- a/src/algebra/modeq.lean +++ b/src/algebra/modeq.lean @@ -9,6 +9,9 @@ import group_theory.quotient_group /-! # Equality modulo an element +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines equality modulo an element in a commutative group. ## Main definitions diff --git a/src/algebra/ring/boolean_ring.lean b/src/algebra/ring/boolean_ring.lean index f93c74cd170fc..2e892c98278e8 100644 --- a/src/algebra/ring/boolean_ring.lean +++ b/src/algebra/ring/boolean_ring.lean @@ -11,6 +11,9 @@ import order.hom.lattice /-! # Boolean rings +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A Boolean ring is a ring where multiplication is idempotent. They are equivalent to Boolean algebras. diff --git a/src/analysis/locally_convex/continuous_of_bounded.lean b/src/analysis/locally_convex/continuous_of_bounded.lean index 7245d211bd25b..d21f6381b2046 100644 --- a/src/analysis/locally_convex/continuous_of_bounded.lean +++ b/src/analysis/locally_convex/continuous_of_bounded.lean @@ -9,6 +9,9 @@ import data.is_R_or_C.basic /-! # Continuity and Von Neumann boundedness +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This files proves that for `E` and `F` two topological vector spaces over `ℝ` or `ℂ`, if `E` is first countable, then every locally bounded linear map `E →ₛₗ[σ] F` is continuous (this is `linear_map.continuous_of_locally_bounded`). diff --git a/src/analysis/normed_space/units.lean b/src/analysis/normed_space/units.lean index 5120eae715bb4..c42e8fcb34961 100644 --- a/src/analysis/normed_space/units.lean +++ b/src/analysis/normed_space/units.lean @@ -9,6 +9,9 @@ import analysis.specific_limits.normed /-! # The group of units of a complete normed ring +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains the basic theory for the group of units (invertible elements) of a complete normed ring (Banach algebras being a notable special case). diff --git a/src/category_theory/abelian/diagram_lemmas/four.lean b/src/category_theory/abelian/diagram_lemmas/four.lean index 6bb067fc43cb0..5cd6428734fc1 100644 --- a/src/category_theory/abelian/diagram_lemmas/four.lean +++ b/src/category_theory/abelian/diagram_lemmas/four.lean @@ -8,6 +8,9 @@ import category_theory.abelian.pseudoelements /-! # The four and five lemmas +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Consider the following commutative diagram with exact rows in an abelian category: ``` diff --git a/src/category_theory/abelian/generator.lean b/src/category_theory/abelian/generator.lean index ef01a1539af08..0b943135e7aa6 100644 --- a/src/category_theory/abelian/generator.lean +++ b/src/category_theory/abelian/generator.lean @@ -12,6 +12,9 @@ import category_theory.abelian.opposite /-! # A complete abelian category with enough injectives and a separator has an injective coseparator +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Future work * Once we know that Grothendieck categories have enough injectives, we can use this to conclude that Grothendieck categories have an injective coseparator. diff --git a/src/category_theory/idempotents/biproducts.lean b/src/category_theory/idempotents/biproducts.lean index 0041beec748fe..88c8757c3259c 100644 --- a/src/category_theory/idempotents/biproducts.lean +++ b/src/category_theory/idempotents/biproducts.lean @@ -10,6 +10,9 @@ import category_theory.idempotents.karoubi # Biproducts in the idempotent completion of a preadditive category +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we define an instance expressing that if `C` is an additive category (i.e. is preadditive and has finite biproducts), then `karoubi C` is also an additive category. diff --git a/src/data/is_R_or_C/basic.lean b/src/data/is_R_or_C/basic.lean index d3907aa627b28..c2db17c4e3424 100644 --- a/src/data/is_R_or_C/basic.lean +++ b/src/data/is_R_or_C/basic.lean @@ -10,6 +10,9 @@ import analysis.normed_space.continuous_linear_map /-! # `is_R_or_C`: a typeclass for ℝ or ℂ +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the typeclass `is_R_or_C` intended to have only two instances: ℝ and ℂ. It is meant for definitions and theorems which hold for both the real and the complex case, and in particular when the real case follows directly from the complex case by setting `re` to `id`, diff --git a/src/model_theory/definability.lean b/src/model_theory/definability.lean index e689a61f60d3a..b38b5f355e3ba 100644 --- a/src/model_theory/definability.lean +++ b/src/model_theory/definability.lean @@ -8,6 +8,9 @@ import model_theory.semantics /-! # Definable Sets + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. This file defines what it means for a set over a first-order structure to be definable. ## Main Definitions diff --git a/src/model_theory/order.lean b/src/model_theory/order.lean index c14ca9cadf61d..d610e1d92fa01 100644 --- a/src/model_theory/order.lean +++ b/src/model_theory/order.lean @@ -7,6 +7,9 @@ import model_theory.semantics /-! # Ordered First-Ordered Structures + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. This file defines ordered first-order languages and structures, as well as their theories. ## Main Definitions diff --git a/src/model_theory/semantics.lean b/src/model_theory/semantics.lean index 724ef54eebf99..a305b8a4449d8 100644 --- a/src/model_theory/semantics.lean +++ b/src/model_theory/semantics.lean @@ -8,6 +8,9 @@ import model_theory.syntax /-! # Basics on First-Order Semantics + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. This file defines the interpretations of first-order terms, formulas, sentences, and theories in a style inspired by the [Flypitch project](https://flypitch.github.io/). diff --git a/src/model_theory/substructures.lean b/src/model_theory/substructures.lean index 8228ffe11dfbf..bd98b13c40d2c 100644 --- a/src/model_theory/substructures.lean +++ b/src/model_theory/substructures.lean @@ -10,6 +10,9 @@ import model_theory.encoding /-! # First-Order Substructures + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. This file defines substructures of first-order structures in a similar manner to the various substructures appearing in the algebra library. diff --git a/src/topology/category/Top/limits/pullbacks.lean b/src/topology/category/Top/limits/pullbacks.lean index 56362bd553478..8e0d847277388 100644 --- a/src/topology/category/Top/limits/pullbacks.lean +++ b/src/topology/category/Top/limits/pullbacks.lean @@ -9,6 +9,9 @@ import category_theory.concrete_category.elementwise /-! # Pullbacks in the category of topological spaces. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + -/ open topological_space From e3fb84046afd187b710170887195d50bada934ee Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Sat, 13 May 2023 18:07:21 +0000 Subject: [PATCH 1004/1291] chore(analysis/calculus/fderiv): split huge file into many pieces (#18995) Splits `analysis/calculus/fderiv.lean` into many subfiles. The file was already structured into well-organized sections. Small sections that fit well-together contentwise end up in a single file. Larger sections get their own file. The order of the sections changed a bit (to the extent that it still makes sense to say that they are ordered, after the split). But none of the sections were split into subsections. Co-authored-by: Eric Wieser --- .../box_integral/divergence_theorem.lean | 5 +- .../calculus/conformal/normed_space.lean | 5 +- src/analysis/calculus/cont_diff_def.lean | 5 +- src/analysis/calculus/deriv.lean | 5 +- src/analysis/calculus/fderiv/add.lean | 572 +++++ src/analysis/calculus/fderiv/basic.lean | 2254 +---------------- src/analysis/calculus/fderiv/bilinear.lean | 139 + src/analysis/calculus/fderiv/comp.lean | 243 ++ src/analysis/calculus/fderiv/equiv.lean | 504 ++++ src/analysis/calculus/fderiv/linear.lean | 124 + src/analysis/calculus/fderiv/mul.lean | 465 ++++ src/analysis/calculus/fderiv/prod.lean | 383 +++ .../calculus/fderiv/restrict_scalars.lean | 101 + 13 files changed, 2562 insertions(+), 2243 deletions(-) create mode 100644 src/analysis/calculus/fderiv/add.lean create mode 100644 src/analysis/calculus/fderiv/bilinear.lean create mode 100644 src/analysis/calculus/fderiv/comp.lean create mode 100644 src/analysis/calculus/fderiv/equiv.lean create mode 100644 src/analysis/calculus/fderiv/linear.lean create mode 100644 src/analysis/calculus/fderiv/mul.lean create mode 100644 src/analysis/calculus/fderiv/prod.lean create mode 100644 src/analysis/calculus/fderiv/restrict_scalars.lean diff --git a/src/analysis/box_integral/divergence_theorem.lean b/src/analysis/box_integral/divergence_theorem.lean index ea65ee8a4ed60..14cdbba22fc83 100644 --- a/src/analysis/box_integral/divergence_theorem.lean +++ b/src/analysis/box_integral/divergence_theorem.lean @@ -5,7 +5,10 @@ Authors: Yury Kudryashov -/ import analysis.box_integral.basic import analysis.box_integral.partition.additive -import analysis.calculus.fderiv.basic +import analysis.calculus.fderiv.add +import analysis.calculus.fderiv.mul +import analysis.calculus.fderiv.equiv +import analysis.calculus.fderiv.restrict_scalars /-! # Divergence integral for Henstock-Kurzweil integral diff --git a/src/analysis/calculus/conformal/normed_space.lean b/src/analysis/calculus/conformal/normed_space.lean index a187e8103b8b7..04916db1c4a99 100644 --- a/src/analysis/calculus/conformal/normed_space.lean +++ b/src/analysis/calculus/conformal/normed_space.lean @@ -4,7 +4,10 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yourong Zang -/ import analysis.normed_space.conformal_linear_map -import analysis.calculus.fderiv.basic +import analysis.calculus.fderiv.add +import analysis.calculus.fderiv.mul +import analysis.calculus.fderiv.equiv +import analysis.calculus.fderiv.restrict_scalars /-! # Conformal Maps diff --git a/src/analysis/calculus/cont_diff_def.lean b/src/analysis/calculus/cont_diff_def.lean index 51293f418f5bc..8e64de8d50d92 100644 --- a/src/analysis/calculus/cont_diff_def.lean +++ b/src/analysis/calculus/cont_diff_def.lean @@ -3,7 +3,10 @@ Copyright (c) 2019 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ -import analysis.calculus.fderiv.basic +import analysis.calculus.fderiv.add +import analysis.calculus.fderiv.mul +import analysis.calculus.fderiv.equiv +import analysis.calculus.fderiv.restrict_scalars import analysis.calculus.formal_multilinear_series /-! diff --git a/src/analysis/calculus/deriv.lean b/src/analysis/calculus/deriv.lean index 4b673fbaf198b..178b32111773e 100644 --- a/src/analysis/calculus/deriv.lean +++ b/src/analysis/calculus/deriv.lean @@ -3,7 +3,10 @@ Copyright (c) 2019 Gabriel Ebner. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Gabriel Ebner, Sébastien Gouëzel -/ -import analysis.calculus.fderiv.basic +import analysis.calculus.fderiv.add +import analysis.calculus.fderiv.mul +import analysis.calculus.fderiv.equiv +import analysis.calculus.fderiv.restrict_scalars import data.polynomial.algebra_map import data.polynomial.derivative import linear_algebra.affine_space.slope diff --git a/src/analysis/calculus/fderiv/add.lean b/src/analysis/calculus/fderiv/add.lean new file mode 100644 index 0000000000000..98c66b263e07c --- /dev/null +++ b/src/analysis/calculus/fderiv/add.lean @@ -0,0 +1,572 @@ +/- +Copyright (c) 2019 Jeremy Avigad. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jeremy Avigad, Sébastien Gouëzel, Yury Kudryashov +-/ +import analysis.calculus.fderiv.linear +import analysis.calculus.fderiv.comp + +/-! +# Additive operations on derivatives + +For detailed documentation of the Fréchet derivative, +see the module docstring of `analysis/calculus/fderiv/basic.lean`. + +This file contains the usual formulas (and existence assertions) for the derivative of + +* sum of finitely many functions +* multiplication of a function by a scalar constant +* negative of a function +* subtraction of two functions +-/ + +open filter asymptotics continuous_linear_map set metric +open_locale topology classical nnreal filter asymptotics ennreal + +noncomputable theory + + +section + +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] +variables {E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] +variables {F : Type*} [normed_add_comm_group F] [normed_space 𝕜 F] +variables {G : Type*} [normed_add_comm_group G] [normed_space 𝕜 G] +variables {G' : Type*} [normed_add_comm_group G'] [normed_space 𝕜 G'] + +variables {f f₀ f₁ g : E → F} +variables {f' f₀' f₁' g' : E →L[𝕜] F} +variables (e : E →L[𝕜] F) +variables {x : E} +variables {s t : set E} +variables {L L₁ L₂ : filter E} + +section const_smul + +variables {R : Type*} [semiring R] [module R F] [smul_comm_class 𝕜 R F] + [has_continuous_const_smul R F] + +/-! ### Derivative of a function multiplied by a constant -/ +theorem has_strict_fderiv_at.const_smul (h : has_strict_fderiv_at f f' x) (c : R) : + has_strict_fderiv_at (λ x, c • f x) (c • f') x := +(c • (1 : F →L[𝕜] F)).has_strict_fderiv_at.comp x h + +theorem has_fderiv_at_filter.const_smul (h : has_fderiv_at_filter f f' x L) (c : R) : + has_fderiv_at_filter (λ x, c • f x) (c • f') x L := +(c • (1 : F →L[𝕜] F)).has_fderiv_at_filter.comp x h tendsto_map + +theorem has_fderiv_within_at.const_smul (h : has_fderiv_within_at f f' s x) (c : R) : + has_fderiv_within_at (λ x, c • f x) (c • f') s x := +h.const_smul c + +theorem has_fderiv_at.const_smul (h : has_fderiv_at f f' x) (c : R) : + has_fderiv_at (λ x, c • f x) (c • f') x := +h.const_smul c + +lemma differentiable_within_at.const_smul (h : differentiable_within_at 𝕜 f s x) (c : R) : + differentiable_within_at 𝕜 (λy, c • f y) s x := +(h.has_fderiv_within_at.const_smul c).differentiable_within_at + +lemma differentiable_at.const_smul (h : differentiable_at 𝕜 f x) (c : R) : + differentiable_at 𝕜 (λy, c • f y) x := +(h.has_fderiv_at.const_smul c).differentiable_at + +lemma differentiable_on.const_smul (h : differentiable_on 𝕜 f s) (c : R) : + differentiable_on 𝕜 (λy, c • f y) s := +λx hx, (h x hx).const_smul c + +lemma differentiable.const_smul (h : differentiable 𝕜 f) (c : R) : + differentiable 𝕜 (λy, c • f y) := +λx, (h x).const_smul c + +lemma fderiv_within_const_smul (hxs : unique_diff_within_at 𝕜 s x) + (h : differentiable_within_at 𝕜 f s x) (c : R) : + fderiv_within 𝕜 (λy, c • f y) s x = c • fderiv_within 𝕜 f s x := +(h.has_fderiv_within_at.const_smul c).fderiv_within hxs + +lemma fderiv_const_smul (h : differentiable_at 𝕜 f x) (c : R) : + fderiv 𝕜 (λy, c • f y) x = c • fderiv 𝕜 f x := +(h.has_fderiv_at.const_smul c).fderiv + +end const_smul + +section add + +/-! ### Derivative of the sum of two functions -/ + +theorem has_strict_fderiv_at.add (hf : has_strict_fderiv_at f f' x) + (hg : has_strict_fderiv_at g g' x) : + has_strict_fderiv_at (λ y, f y + g y) (f' + g') x := +(hf.add hg).congr_left $ λ y, + by { simp only [linear_map.sub_apply, linear_map.add_apply, map_sub, map_add, add_apply], abel } + +theorem has_fderiv_at_filter.add + (hf : has_fderiv_at_filter f f' x L) (hg : has_fderiv_at_filter g g' x L) : + has_fderiv_at_filter (λ y, f y + g y) (f' + g') x L := +(hf.add hg).congr_left $ λ _, + by { simp only [linear_map.sub_apply, linear_map.add_apply, map_sub, map_add, add_apply], abel } + +theorem has_fderiv_within_at.add + (hf : has_fderiv_within_at f f' s x) (hg : has_fderiv_within_at g g' s x) : + has_fderiv_within_at (λ y, f y + g y) (f' + g') s x := +hf.add hg + +theorem has_fderiv_at.add + (hf : has_fderiv_at f f' x) (hg : has_fderiv_at g g' x) : + has_fderiv_at (λ x, f x + g x) (f' + g') x := +hf.add hg + +lemma differentiable_within_at.add + (hf : differentiable_within_at 𝕜 f s x) (hg : differentiable_within_at 𝕜 g s x) : + differentiable_within_at 𝕜 (λ y, f y + g y) s x := +(hf.has_fderiv_within_at.add hg.has_fderiv_within_at).differentiable_within_at + +@[simp] lemma differentiable_at.add + (hf : differentiable_at 𝕜 f x) (hg : differentiable_at 𝕜 g x) : + differentiable_at 𝕜 (λ y, f y + g y) x := +(hf.has_fderiv_at.add hg.has_fderiv_at).differentiable_at + +lemma differentiable_on.add + (hf : differentiable_on 𝕜 f s) (hg : differentiable_on 𝕜 g s) : + differentiable_on 𝕜 (λy, f y + g y) s := +λx hx, (hf x hx).add (hg x hx) + +@[simp] lemma differentiable.add + (hf : differentiable 𝕜 f) (hg : differentiable 𝕜 g) : + differentiable 𝕜 (λy, f y + g y) := +λx, (hf x).add (hg x) + +lemma fderiv_within_add (hxs : unique_diff_within_at 𝕜 s x) + (hf : differentiable_within_at 𝕜 f s x) (hg : differentiable_within_at 𝕜 g s x) : + fderiv_within 𝕜 (λy, f y + g y) s x = fderiv_within 𝕜 f s x + fderiv_within 𝕜 g s x := +(hf.has_fderiv_within_at.add hg.has_fderiv_within_at).fderiv_within hxs + +lemma fderiv_add + (hf : differentiable_at 𝕜 f x) (hg : differentiable_at 𝕜 g x) : + fderiv 𝕜 (λy, f y + g y) x = fderiv 𝕜 f x + fderiv 𝕜 g x := +(hf.has_fderiv_at.add hg.has_fderiv_at).fderiv + +theorem has_strict_fderiv_at.add_const (hf : has_strict_fderiv_at f f' x) (c : F) : + has_strict_fderiv_at (λ y, f y + c) f' x := +add_zero f' ▸ hf.add (has_strict_fderiv_at_const _ _) + +theorem has_fderiv_at_filter.add_const + (hf : has_fderiv_at_filter f f' x L) (c : F) : + has_fderiv_at_filter (λ y, f y + c) f' x L := +add_zero f' ▸ hf.add (has_fderiv_at_filter_const _ _ _) + +theorem has_fderiv_within_at.add_const + (hf : has_fderiv_within_at f f' s x) (c : F) : + has_fderiv_within_at (λ y, f y + c) f' s x := +hf.add_const c + +theorem has_fderiv_at.add_const (hf : has_fderiv_at f f' x) (c : F): + has_fderiv_at (λ x, f x + c) f' x := +hf.add_const c + +lemma differentiable_within_at.add_const + (hf : differentiable_within_at 𝕜 f s x) (c : F) : + differentiable_within_at 𝕜 (λ y, f y + c) s x := +(hf.has_fderiv_within_at.add_const c).differentiable_within_at + +@[simp] lemma differentiable_within_at_add_const_iff (c : F) : + differentiable_within_at 𝕜 (λ y, f y + c) s x ↔ differentiable_within_at 𝕜 f s x := +⟨λ h, by simpa using h.add_const (-c), λ h, h.add_const c⟩ + +lemma differentiable_at.add_const + (hf : differentiable_at 𝕜 f x) (c : F) : + differentiable_at 𝕜 (λ y, f y + c) x := +(hf.has_fderiv_at.add_const c).differentiable_at + +@[simp] lemma differentiable_at_add_const_iff (c : F) : + differentiable_at 𝕜 (λ y, f y + c) x ↔ differentiable_at 𝕜 f x := +⟨λ h, by simpa using h.add_const (-c), λ h, h.add_const c⟩ + +lemma differentiable_on.add_const + (hf : differentiable_on 𝕜 f s) (c : F) : + differentiable_on 𝕜 (λy, f y + c) s := +λx hx, (hf x hx).add_const c + +@[simp] lemma differentiable_on_add_const_iff (c : F) : + differentiable_on 𝕜 (λ y, f y + c) s ↔ differentiable_on 𝕜 f s := +⟨λ h, by simpa using h.add_const (-c), λ h, h.add_const c⟩ + +lemma differentiable.add_const + (hf : differentiable 𝕜 f) (c : F) : + differentiable 𝕜 (λy, f y + c) := +λx, (hf x).add_const c + +@[simp] lemma differentiable_add_const_iff (c : F) : + differentiable 𝕜 (λ y, f y + c) ↔ differentiable 𝕜 f := +⟨λ h, by simpa using h.add_const (-c), λ h, h.add_const c⟩ + +lemma fderiv_within_add_const (hxs : unique_diff_within_at 𝕜 s x) (c : F) : + fderiv_within 𝕜 (λy, f y + c) s x = fderiv_within 𝕜 f s x := +if hf : differentiable_within_at 𝕜 f s x +then (hf.has_fderiv_within_at.add_const c).fderiv_within hxs +else by { rw [fderiv_within_zero_of_not_differentiable_within_at hf, + fderiv_within_zero_of_not_differentiable_within_at], simpa } + +lemma fderiv_add_const (c : F) : fderiv 𝕜 (λy, f y + c) x = fderiv 𝕜 f x := +by simp only [← fderiv_within_univ, fderiv_within_add_const unique_diff_within_at_univ] + +theorem has_strict_fderiv_at.const_add (hf : has_strict_fderiv_at f f' x) (c : F) : + has_strict_fderiv_at (λ y, c + f y) f' x := +zero_add f' ▸ (has_strict_fderiv_at_const _ _).add hf + +theorem has_fderiv_at_filter.const_add + (hf : has_fderiv_at_filter f f' x L) (c : F) : + has_fderiv_at_filter (λ y, c + f y) f' x L := +zero_add f' ▸ (has_fderiv_at_filter_const _ _ _).add hf + +theorem has_fderiv_within_at.const_add + (hf : has_fderiv_within_at f f' s x) (c : F) : + has_fderiv_within_at (λ y, c + f y) f' s x := +hf.const_add c + +theorem has_fderiv_at.const_add + (hf : has_fderiv_at f f' x) (c : F): + has_fderiv_at (λ x, c + f x) f' x := +hf.const_add c + +lemma differentiable_within_at.const_add + (hf : differentiable_within_at 𝕜 f s x) (c : F) : + differentiable_within_at 𝕜 (λ y, c + f y) s x := +(hf.has_fderiv_within_at.const_add c).differentiable_within_at + +@[simp] lemma differentiable_within_at_const_add_iff (c : F) : + differentiable_within_at 𝕜 (λ y, c + f y) s x ↔ differentiable_within_at 𝕜 f s x := +⟨λ h, by simpa using h.const_add (-c), λ h, h.const_add c⟩ + +lemma differentiable_at.const_add + (hf : differentiable_at 𝕜 f x) (c : F) : + differentiable_at 𝕜 (λ y, c + f y) x := +(hf.has_fderiv_at.const_add c).differentiable_at + +@[simp] lemma differentiable_at_const_add_iff (c : F) : + differentiable_at 𝕜 (λ y, c + f y) x ↔ differentiable_at 𝕜 f x := +⟨λ h, by simpa using h.const_add (-c), λ h, h.const_add c⟩ + +lemma differentiable_on.const_add (hf : differentiable_on 𝕜 f s) (c : F) : + differentiable_on 𝕜 (λy, c + f y) s := +λx hx, (hf x hx).const_add c + +@[simp] lemma differentiable_on_const_add_iff (c : F) : + differentiable_on 𝕜 (λ y, c + f y) s ↔ differentiable_on 𝕜 f s := +⟨λ h, by simpa using h.const_add (-c), λ h, h.const_add c⟩ + +lemma differentiable.const_add (hf : differentiable 𝕜 f) (c : F) : + differentiable 𝕜 (λy, c + f y) := +λx, (hf x).const_add c + +@[simp] lemma differentiable_const_add_iff (c : F) : + differentiable 𝕜 (λ y, c + f y) ↔ differentiable 𝕜 f := +⟨λ h, by simpa using h.const_add (-c), λ h, h.const_add c⟩ + +lemma fderiv_within_const_add (hxs : unique_diff_within_at 𝕜 s x) (c : F) : + fderiv_within 𝕜 (λy, c + f y) s x = fderiv_within 𝕜 f s x := +by simpa only [add_comm] using fderiv_within_add_const hxs c + +lemma fderiv_const_add (c : F) : fderiv 𝕜 (λy, c + f y) x = fderiv 𝕜 f x := +by simp only [add_comm c, fderiv_add_const] + +end add + +section sum +/-! ### Derivative of a finite sum of functions -/ + +open_locale big_operators + +variables {ι : Type*} {u : finset ι} {A : ι → (E → F)} {A' : ι → (E →L[𝕜] F)} + +theorem has_strict_fderiv_at.sum (h : ∀ i ∈ u, has_strict_fderiv_at (A i) (A' i) x) : + has_strict_fderiv_at (λ y, ∑ i in u, A i y) (∑ i in u, A' i) x := +begin + dsimp [has_strict_fderiv_at] at *, + convert is_o.sum h, + simp [finset.sum_sub_distrib, continuous_linear_map.sum_apply] +end + +theorem has_fderiv_at_filter.sum (h : ∀ i ∈ u, has_fderiv_at_filter (A i) (A' i) x L) : + has_fderiv_at_filter (λ y, ∑ i in u, A i y) (∑ i in u, A' i) x L := +begin + dsimp [has_fderiv_at_filter] at *, + convert is_o.sum h, + simp [continuous_linear_map.sum_apply] +end + +theorem has_fderiv_within_at.sum (h : ∀ i ∈ u, has_fderiv_within_at (A i) (A' i) s x) : + has_fderiv_within_at (λ y, ∑ i in u, A i y) (∑ i in u, A' i) s x := +has_fderiv_at_filter.sum h + +theorem has_fderiv_at.sum (h : ∀ i ∈ u, has_fderiv_at (A i) (A' i) x) : + has_fderiv_at (λ y, ∑ i in u, A i y) (∑ i in u, A' i) x := +has_fderiv_at_filter.sum h + +theorem differentiable_within_at.sum (h : ∀ i ∈ u, differentiable_within_at 𝕜 (A i) s x) : + differentiable_within_at 𝕜 (λ y, ∑ i in u, A i y) s x := +has_fderiv_within_at.differentiable_within_at $ has_fderiv_within_at.sum $ +λ i hi, (h i hi).has_fderiv_within_at + +@[simp] theorem differentiable_at.sum (h : ∀ i ∈ u, differentiable_at 𝕜 (A i) x) : + differentiable_at 𝕜 (λ y, ∑ i in u, A i y) x := +has_fderiv_at.differentiable_at $ has_fderiv_at.sum $ λ i hi, (h i hi).has_fderiv_at + +theorem differentiable_on.sum (h : ∀ i ∈ u, differentiable_on 𝕜 (A i) s) : + differentiable_on 𝕜 (λ y, ∑ i in u, A i y) s := +λ x hx, differentiable_within_at.sum $ λ i hi, h i hi x hx + +@[simp] theorem differentiable.sum (h : ∀ i ∈ u, differentiable 𝕜 (A i)) : + differentiable 𝕜 (λ y, ∑ i in u, A i y) := +λ x, differentiable_at.sum $ λ i hi, h i hi x + +theorem fderiv_within_sum (hxs : unique_diff_within_at 𝕜 s x) + (h : ∀ i ∈ u, differentiable_within_at 𝕜 (A i) s x) : + fderiv_within 𝕜 (λ y, ∑ i in u, A i y) s x = (∑ i in u, fderiv_within 𝕜 (A i) s x) := +(has_fderiv_within_at.sum (λ i hi, (h i hi).has_fderiv_within_at)).fderiv_within hxs + +theorem fderiv_sum (h : ∀ i ∈ u, differentiable_at 𝕜 (A i) x) : + fderiv 𝕜 (λ y, ∑ i in u, A i y) x = (∑ i in u, fderiv 𝕜 (A i) x) := +(has_fderiv_at.sum (λ i hi, (h i hi).has_fderiv_at)).fderiv + +end sum + +section neg +/-! ### Derivative of the negative of a function -/ + +theorem has_strict_fderiv_at.neg (h : has_strict_fderiv_at f f' x) : + has_strict_fderiv_at (λ x, -f x) (-f') x := +(-1 : F →L[𝕜] F).has_strict_fderiv_at.comp x h + +theorem has_fderiv_at_filter.neg (h : has_fderiv_at_filter f f' x L) : + has_fderiv_at_filter (λ x, -f x) (-f') x L := +(-1 : F →L[𝕜] F).has_fderiv_at_filter.comp x h tendsto_map + +theorem has_fderiv_within_at.neg (h : has_fderiv_within_at f f' s x) : + has_fderiv_within_at (λ x, -f x) (-f') s x := +h.neg + +theorem has_fderiv_at.neg (h : has_fderiv_at f f' x) : + has_fderiv_at (λ x, -f x) (-f') x := +h.neg + +lemma differentiable_within_at.neg (h : differentiable_within_at 𝕜 f s x) : + differentiable_within_at 𝕜 (λy, -f y) s x := +h.has_fderiv_within_at.neg.differentiable_within_at + +@[simp] lemma differentiable_within_at_neg_iff : + differentiable_within_at 𝕜 (λy, -f y) s x ↔ differentiable_within_at 𝕜 f s x := +⟨λ h, by simpa only [neg_neg] using h.neg, λ h, h.neg⟩ + +lemma differentiable_at.neg (h : differentiable_at 𝕜 f x) : + differentiable_at 𝕜 (λy, -f y) x := +h.has_fderiv_at.neg.differentiable_at + +@[simp] lemma differentiable_at_neg_iff : + differentiable_at 𝕜 (λy, -f y) x ↔ differentiable_at 𝕜 f x := +⟨λ h, by simpa only [neg_neg] using h.neg, λ h, h.neg⟩ + +lemma differentiable_on.neg (h : differentiable_on 𝕜 f s) : + differentiable_on 𝕜 (λy, -f y) s := +λx hx, (h x hx).neg + +@[simp] lemma differentiable_on_neg_iff : + differentiable_on 𝕜 (λy, -f y) s ↔ differentiable_on 𝕜 f s := +⟨λ h, by simpa only [neg_neg] using h.neg, λ h, h.neg⟩ + +lemma differentiable.neg (h : differentiable 𝕜 f) : + differentiable 𝕜 (λy, -f y) := +λx, (h x).neg + +@[simp] lemma differentiable_neg_iff : differentiable 𝕜 (λy, -f y) ↔ differentiable 𝕜 f := +⟨λ h, by simpa only [neg_neg] using h.neg, λ h, h.neg⟩ + +lemma fderiv_within_neg (hxs : unique_diff_within_at 𝕜 s x) : + fderiv_within 𝕜 (λy, -f y) s x = - fderiv_within 𝕜 f s x := +if h : differentiable_within_at 𝕜 f s x +then h.has_fderiv_within_at.neg.fderiv_within hxs +else by { rw [fderiv_within_zero_of_not_differentiable_within_at h, + fderiv_within_zero_of_not_differentiable_within_at, neg_zero], simpa } + +@[simp] lemma fderiv_neg : fderiv 𝕜 (λy, -f y) x = - fderiv 𝕜 f x := +by simp only [← fderiv_within_univ, fderiv_within_neg unique_diff_within_at_univ] + +end neg + +section sub +/-! ### Derivative of the difference of two functions -/ + +theorem has_strict_fderiv_at.sub + (hf : has_strict_fderiv_at f f' x) (hg : has_strict_fderiv_at g g' x) : + has_strict_fderiv_at (λ x, f x - g x) (f' - g') x := +by simpa only [sub_eq_add_neg] using hf.add hg.neg + +theorem has_fderiv_at_filter.sub + (hf : has_fderiv_at_filter f f' x L) (hg : has_fderiv_at_filter g g' x L) : + has_fderiv_at_filter (λ x, f x - g x) (f' - g') x L := +by simpa only [sub_eq_add_neg] using hf.add hg.neg + +theorem has_fderiv_within_at.sub + (hf : has_fderiv_within_at f f' s x) (hg : has_fderiv_within_at g g' s x) : + has_fderiv_within_at (λ x, f x - g x) (f' - g') s x := +hf.sub hg + +theorem has_fderiv_at.sub + (hf : has_fderiv_at f f' x) (hg : has_fderiv_at g g' x) : + has_fderiv_at (λ x, f x - g x) (f' - g') x := +hf.sub hg + +lemma differentiable_within_at.sub + (hf : differentiable_within_at 𝕜 f s x) (hg : differentiable_within_at 𝕜 g s x) : + differentiable_within_at 𝕜 (λ y, f y - g y) s x := +(hf.has_fderiv_within_at.sub hg.has_fderiv_within_at).differentiable_within_at + +@[simp] lemma differentiable_at.sub + (hf : differentiable_at 𝕜 f x) (hg : differentiable_at 𝕜 g x) : + differentiable_at 𝕜 (λ y, f y - g y) x := +(hf.has_fderiv_at.sub hg.has_fderiv_at).differentiable_at + +lemma differentiable_on.sub + (hf : differentiable_on 𝕜 f s) (hg : differentiable_on 𝕜 g s) : + differentiable_on 𝕜 (λy, f y - g y) s := +λx hx, (hf x hx).sub (hg x hx) + +@[simp] lemma differentiable.sub + (hf : differentiable 𝕜 f) (hg : differentiable 𝕜 g) : + differentiable 𝕜 (λy, f y - g y) := +λx, (hf x).sub (hg x) + +lemma fderiv_within_sub (hxs : unique_diff_within_at 𝕜 s x) + (hf : differentiable_within_at 𝕜 f s x) (hg : differentiable_within_at 𝕜 g s x) : + fderiv_within 𝕜 (λy, f y - g y) s x = fderiv_within 𝕜 f s x - fderiv_within 𝕜 g s x := +(hf.has_fderiv_within_at.sub hg.has_fderiv_within_at).fderiv_within hxs + +lemma fderiv_sub + (hf : differentiable_at 𝕜 f x) (hg : differentiable_at 𝕜 g x) : + fderiv 𝕜 (λy, f y - g y) x = fderiv 𝕜 f x - fderiv 𝕜 g x := +(hf.has_fderiv_at.sub hg.has_fderiv_at).fderiv + +theorem has_strict_fderiv_at.sub_const + (hf : has_strict_fderiv_at f f' x) (c : F) : + has_strict_fderiv_at (λ x, f x - c) f' x := +by simpa only [sub_eq_add_neg] using hf.add_const (-c) + +theorem has_fderiv_at_filter.sub_const + (hf : has_fderiv_at_filter f f' x L) (c : F) : + has_fderiv_at_filter (λ x, f x - c) f' x L := +by simpa only [sub_eq_add_neg] using hf.add_const (-c) + +theorem has_fderiv_within_at.sub_const + (hf : has_fderiv_within_at f f' s x) (c : F) : + has_fderiv_within_at (λ x, f x - c) f' s x := +hf.sub_const c + +theorem has_fderiv_at.sub_const + (hf : has_fderiv_at f f' x) (c : F) : + has_fderiv_at (λ x, f x - c) f' x := +hf.sub_const c + +lemma differentiable_within_at.sub_const + (hf : differentiable_within_at 𝕜 f s x) (c : F) : + differentiable_within_at 𝕜 (λ y, f y - c) s x := +(hf.has_fderiv_within_at.sub_const c).differentiable_within_at + +@[simp] lemma differentiable_within_at_sub_const_iff (c : F) : + differentiable_within_at 𝕜 (λ y, f y - c) s x ↔ differentiable_within_at 𝕜 f s x := +by simp only [sub_eq_add_neg, differentiable_within_at_add_const_iff] + +lemma differentiable_at.sub_const (hf : differentiable_at 𝕜 f x) (c : F) : + differentiable_at 𝕜 (λ y, f y - c) x := +(hf.has_fderiv_at.sub_const c).differentiable_at + +@[simp] lemma differentiable_at_sub_const_iff (c : F) : + differentiable_at 𝕜 (λ y, f y - c) x ↔ differentiable_at 𝕜 f x := +by simp only [sub_eq_add_neg, differentiable_at_add_const_iff] + +lemma differentiable_on.sub_const (hf : differentiable_on 𝕜 f s) (c : F) : + differentiable_on 𝕜 (λy, f y - c) s := +λx hx, (hf x hx).sub_const c + +@[simp] lemma differentiable_on_sub_const_iff (c : F) : + differentiable_on 𝕜 (λ y, f y - c) s ↔ differentiable_on 𝕜 f s := +by simp only [sub_eq_add_neg, differentiable_on_add_const_iff] + +lemma differentiable.sub_const (hf : differentiable 𝕜 f) (c : F) : + differentiable 𝕜 (λy, f y - c) := +λx, (hf x).sub_const c + +@[simp] lemma differentiable_sub_const_iff (c : F) : + differentiable 𝕜 (λ y, f y - c) ↔ differentiable 𝕜 f := +by simp only [sub_eq_add_neg, differentiable_add_const_iff] + +lemma fderiv_within_sub_const (hxs : unique_diff_within_at 𝕜 s x) (c : F) : + fderiv_within 𝕜 (λy, f y - c) s x = fderiv_within 𝕜 f s x := +by simp only [sub_eq_add_neg, fderiv_within_add_const hxs] + +lemma fderiv_sub_const (c : F) : fderiv 𝕜 (λy, f y - c) x = fderiv 𝕜 f x := +by simp only [sub_eq_add_neg, fderiv_add_const] + +theorem has_strict_fderiv_at.const_sub + (hf : has_strict_fderiv_at f f' x) (c : F) : + has_strict_fderiv_at (λ x, c - f x) (-f') x := +by simpa only [sub_eq_add_neg] using hf.neg.const_add c + +theorem has_fderiv_at_filter.const_sub + (hf : has_fderiv_at_filter f f' x L) (c : F) : + has_fderiv_at_filter (λ x, c - f x) (-f') x L := +by simpa only [sub_eq_add_neg] using hf.neg.const_add c + +theorem has_fderiv_within_at.const_sub + (hf : has_fderiv_within_at f f' s x) (c : F) : + has_fderiv_within_at (λ x, c - f x) (-f') s x := +hf.const_sub c + +theorem has_fderiv_at.const_sub + (hf : has_fderiv_at f f' x) (c : F) : + has_fderiv_at (λ x, c - f x) (-f') x := +hf.const_sub c + +lemma differentiable_within_at.const_sub + (hf : differentiable_within_at 𝕜 f s x) (c : F) : + differentiable_within_at 𝕜 (λ y, c - f y) s x := +(hf.has_fderiv_within_at.const_sub c).differentiable_within_at + +@[simp] lemma differentiable_within_at_const_sub_iff (c : F) : + differentiable_within_at 𝕜 (λ y, c - f y) s x ↔ differentiable_within_at 𝕜 f s x := +by simp [sub_eq_add_neg] + +lemma differentiable_at.const_sub + (hf : differentiable_at 𝕜 f x) (c : F) : + differentiable_at 𝕜 (λ y, c - f y) x := +(hf.has_fderiv_at.const_sub c).differentiable_at + +@[simp] lemma differentiable_at_const_sub_iff (c : F) : + differentiable_at 𝕜 (λ y, c - f y) x ↔ differentiable_at 𝕜 f x := +by simp [sub_eq_add_neg] + +lemma differentiable_on.const_sub (hf : differentiable_on 𝕜 f s) (c : F) : + differentiable_on 𝕜 (λy, c - f y) s := +λx hx, (hf x hx).const_sub c + +@[simp] lemma differentiable_on_const_sub_iff (c : F) : + differentiable_on 𝕜 (λ y, c - f y) s ↔ differentiable_on 𝕜 f s := +by simp [sub_eq_add_neg] + +lemma differentiable.const_sub (hf : differentiable 𝕜 f) (c : F) : + differentiable 𝕜 (λy, c - f y) := +λx, (hf x).const_sub c + +@[simp] lemma differentiable_const_sub_iff (c : F) : + differentiable 𝕜 (λ y, c - f y) ↔ differentiable 𝕜 f := +by simp [sub_eq_add_neg] + +lemma fderiv_within_const_sub (hxs : unique_diff_within_at 𝕜 s x) (c : F) : + fderiv_within 𝕜 (λy, c - f y) s x = -fderiv_within 𝕜 f s x := +by simp only [sub_eq_add_neg, fderiv_within_const_add, fderiv_within_neg, hxs] + +lemma fderiv_const_sub (c : F) : fderiv 𝕜 (λy, c - f y) x = -fderiv 𝕜 f x := +by simp only [← fderiv_within_univ, fderiv_within_const_sub unique_diff_within_at_univ] + +end sub + +end diff --git a/src/analysis/calculus/fderiv/basic.lean b/src/analysis/calculus/fderiv/basic.lean index de89c70222e52..f97d738340d5f 100644 --- a/src/analysis/calculus/fderiv/basic.lean +++ b/src/analysis/calculus/fderiv/basic.lean @@ -32,21 +32,23 @@ function theorem, and is defined here only to avoid proving theorems like ## Main results -In addition to the definition and basic properties of the derivative, this file contains the -usual formulas (and existence assertions) for the derivative of +In addition to the definition and basic properties of the derivative, +the folder `analysis/calculus/fderiv/` contains the usual formulas +(and existence assertions) for the derivative of * constants * the identity -* bounded linear maps -* bounded bilinear maps -* sum of two functions -* sum of finitely many functions -* multiplication of a function by a scalar constant -* negative of a function -* subtraction of two functions -* multiplication of a function by a scalar function -* multiplication of two scalar functions -* composition of functions (the chain rule) -* inverse function (assuming that it exists; the inverse function theorem is in `inverse.lean`) +* bounded linear maps (`linear.lean`) +* bounded bilinear maps (`bilinear.lean`) +* sum of two functions (`add.lean`) +* sum of finitely many functions (`add.lean`) +* multiplication of a function by a scalar constant (`add.lean`) +* negative of a function (`add.lean`) +* subtraction of two functions (`add.lean`) +* multiplication of a function by a scalar function (`mul.lean`) +* multiplication of two scalar functions (`mul.lean`) +* composition of functions (the chain rule) (`comp.lean`) +* inverse function (`mul.lean`) + (assuming that it exists; the inverse function theorem is in `../inverse.lean`) For most binary operations we also define `const_op` and `op_const` theorems for the cases when the first or second argument is a constant. This makes writing chains of `has_deriv_at`'s easier, @@ -1004,2234 +1006,8 @@ lemma has_fderiv_at_zero_of_eventually_const end const -section continuous_linear_map -/-! -### Continuous linear maps - -There are currently two variants of these in mathlib, the bundled version -(named `continuous_linear_map`, and denoted `E →L[𝕜] F`), and the unbundled version (with a -predicate `is_bounded_linear_map`). We give statements for both versions. -/ - -protected theorem continuous_linear_map.has_strict_fderiv_at {x : E} : - has_strict_fderiv_at e e x := -(is_o_zero _ _).congr_left $ λ x, by simp only [e.map_sub, sub_self] - -protected lemma continuous_linear_map.has_fderiv_at_filter : - has_fderiv_at_filter e e x L := -(is_o_zero _ _).congr_left $ λ x, by simp only [e.map_sub, sub_self] - -protected lemma continuous_linear_map.has_fderiv_within_at : has_fderiv_within_at e e s x := -e.has_fderiv_at_filter - -protected lemma continuous_linear_map.has_fderiv_at : has_fderiv_at e e x := -e.has_fderiv_at_filter - -@[simp] protected lemma continuous_linear_map.differentiable_at : differentiable_at 𝕜 e x := -e.has_fderiv_at.differentiable_at - -protected lemma continuous_linear_map.differentiable_within_at : differentiable_within_at 𝕜 e s x := -e.differentiable_at.differentiable_within_at - -@[simp] protected lemma continuous_linear_map.fderiv : fderiv 𝕜 e x = e := -e.has_fderiv_at.fderiv - -protected lemma continuous_linear_map.fderiv_within (hxs : unique_diff_within_at 𝕜 s x) : - fderiv_within 𝕜 e s x = e := -begin - rw differentiable_at.fderiv_within e.differentiable_at hxs, - exact e.fderiv -end - -@[simp] protected lemma continuous_linear_map.differentiable : differentiable 𝕜 e := -λx, e.differentiable_at - -protected lemma continuous_linear_map.differentiable_on : differentiable_on 𝕜 e s := -e.differentiable.differentiable_on - -lemma is_bounded_linear_map.has_fderiv_at_filter (h : is_bounded_linear_map 𝕜 f) : - has_fderiv_at_filter f h.to_continuous_linear_map x L := -h.to_continuous_linear_map.has_fderiv_at_filter - -lemma is_bounded_linear_map.has_fderiv_within_at (h : is_bounded_linear_map 𝕜 f) : - has_fderiv_within_at f h.to_continuous_linear_map s x := -h.has_fderiv_at_filter - -lemma is_bounded_linear_map.has_fderiv_at (h : is_bounded_linear_map 𝕜 f) : - has_fderiv_at f h.to_continuous_linear_map x := -h.has_fderiv_at_filter - -lemma is_bounded_linear_map.differentiable_at (h : is_bounded_linear_map 𝕜 f) : - differentiable_at 𝕜 f x := -h.has_fderiv_at.differentiable_at - -lemma is_bounded_linear_map.differentiable_within_at (h : is_bounded_linear_map 𝕜 f) : - differentiable_within_at 𝕜 f s x := -h.differentiable_at.differentiable_within_at - -lemma is_bounded_linear_map.fderiv (h : is_bounded_linear_map 𝕜 f) : - fderiv 𝕜 f x = h.to_continuous_linear_map := -has_fderiv_at.fderiv (h.has_fderiv_at) - -lemma is_bounded_linear_map.fderiv_within (h : is_bounded_linear_map 𝕜 f) - (hxs : unique_diff_within_at 𝕜 s x) : fderiv_within 𝕜 f s x = h.to_continuous_linear_map := -begin - rw differentiable_at.fderiv_within h.differentiable_at hxs, - exact h.fderiv -end - -lemma is_bounded_linear_map.differentiable (h : is_bounded_linear_map 𝕜 f) : - differentiable 𝕜 f := -λx, h.differentiable_at - -lemma is_bounded_linear_map.differentiable_on (h : is_bounded_linear_map 𝕜 f) : - differentiable_on 𝕜 f s := -h.differentiable.differentiable_on - -end continuous_linear_map - -section composition -/-! -### Derivative of the composition of two functions - -For composition lemmas, we put x explicit to help the elaborator, as otherwise Lean tends to -get confused since there are too many possibilities for composition -/ - -variable (x) - -theorem has_fderiv_at_filter.comp {g : F → G} {g' : F →L[𝕜] G} {L' : filter F} - (hg : has_fderiv_at_filter g g' (f x) L') - (hf : has_fderiv_at_filter f f' x L) (hL : tendsto f L L') : - has_fderiv_at_filter (g ∘ f) (g'.comp f') x L := -let eq₁ := (g'.is_O_comp _ _).trans_is_o hf in -let eq₂ := (hg.comp_tendsto hL).trans_is_O hf.is_O_sub in -by { refine eq₂.triangle (eq₁.congr_left (λ x', _)), simp } - -/- A readable version of the previous theorem, - a general form of the chain rule. -/ - -example {g : F → G} {g' : F →L[𝕜] G} - (hg : has_fderiv_at_filter g g' (f x) (L.map f)) - (hf : has_fderiv_at_filter f f' x L) : - has_fderiv_at_filter (g ∘ f) (g'.comp f') x L := -begin - unfold has_fderiv_at_filter at hg, - have := calc (λ x', g (f x') - g (f x) - g' (f x' - f x)) =o[L] (λ x', f x' - f x) : - hg.comp_tendsto le_rfl - ... =O[L] (λ x', x' - x) : hf.is_O_sub, - refine this.triangle _, - calc (λ x' : E, g' (f x' - f x) - g'.comp f' (x' - x)) - =ᶠ[L] λ x', g' (f x' - f x - f' (x' - x)) : eventually_of_forall (λ x', by simp) - ... =O[L] λ x', f x' - f x - f' (x' - x) : g'.is_O_comp _ _ - ... =o[L] λ x', x' - x : hf -end - -theorem has_fderiv_within_at.comp {g : F → G} {g' : F →L[𝕜] G} {t : set F} - (hg : has_fderiv_within_at g g' t (f x)) (hf : has_fderiv_within_at f f' s x) - (hst : maps_to f s t) : - has_fderiv_within_at (g ∘ f) (g'.comp f') s x := -hg.comp x hf $ hf.continuous_within_at.tendsto_nhds_within hst - -theorem has_fderiv_at.comp_has_fderiv_within_at {g : F → G} {g' : F →L[𝕜] G} - (hg : has_fderiv_at g g' (f x)) (hf : has_fderiv_within_at f f' s x) : - has_fderiv_within_at (g ∘ f) (g'.comp f') s x := -hg.comp x hf hf.continuous_within_at - -theorem has_fderiv_within_at.comp_of_mem {g : F → G} {g' : F →L[𝕜] G} {t : set F} - (hg : has_fderiv_within_at g g' t (f x)) (hf : has_fderiv_within_at f f' s x) - (hst : tendsto f (𝓝[s] x) (𝓝[t] f x)) : - has_fderiv_within_at (g ∘ f) (g'.comp f') s x := -has_fderiv_at_filter.comp x hg hf hst - -/-- The chain rule. -/ -theorem has_fderiv_at.comp {g : F → G} {g' : F →L[𝕜] G} - (hg : has_fderiv_at g g' (f x)) (hf : has_fderiv_at f f' x) : - has_fderiv_at (g ∘ f) (g'.comp f') x := -hg.comp x hf hf.continuous_at - -lemma differentiable_within_at.comp {g : F → G} {t : set F} - (hg : differentiable_within_at 𝕜 g t (f x)) (hf : differentiable_within_at 𝕜 f s x) - (h : maps_to f s t) : differentiable_within_at 𝕜 (g ∘ f) s x := -(hg.has_fderiv_within_at.comp x hf.has_fderiv_within_at h).differentiable_within_at - -lemma differentiable_within_at.comp' {g : F → G} {t : set F} - (hg : differentiable_within_at 𝕜 g t (f x)) (hf : differentiable_within_at 𝕜 f s x) : - differentiable_within_at 𝕜 (g ∘ f) (s ∩ f⁻¹' t) x := -hg.comp x (hf.mono (inter_subset_left _ _)) (inter_subset_right _ _) - -lemma differentiable_at.comp {g : F → G} - (hg : differentiable_at 𝕜 g (f x)) (hf : differentiable_at 𝕜 f x) : - differentiable_at 𝕜 (g ∘ f) x := -(hg.has_fderiv_at.comp x hf.has_fderiv_at).differentiable_at - -lemma differentiable_at.comp_differentiable_within_at {g : F → G} - (hg : differentiable_at 𝕜 g (f x)) (hf : differentiable_within_at 𝕜 f s x) : - differentiable_within_at 𝕜 (g ∘ f) s x := -hg.differentiable_within_at.comp x hf (maps_to_univ _ _) - -lemma fderiv_within.comp {g : F → G} {t : set F} - (hg : differentiable_within_at 𝕜 g t (f x)) (hf : differentiable_within_at 𝕜 f s x) - (h : maps_to f s t) (hxs : unique_diff_within_at 𝕜 s x) : - fderiv_within 𝕜 (g ∘ f) s x = (fderiv_within 𝕜 g t (f x)).comp (fderiv_within 𝕜 f s x) := -(hg.has_fderiv_within_at.comp x (hf.has_fderiv_within_at) h).fderiv_within hxs - -/-- A version of `fderiv_within.comp` that is useful to rewrite the composition of two derivatives - into a single derivative. This version always applies, but creates a new side-goal `f x = y`. -/ -lemma fderiv_within_fderiv_within {g : F → G} {f : E → F} {x : E} {y : F} {s : set E} {t : set F} - (hg : differentiable_within_at 𝕜 g t y) (hf : differentiable_within_at 𝕜 f s x) - (h : maps_to f s t) (hxs : unique_diff_within_at 𝕜 s x) (hy : f x = y) (v : E) : - fderiv_within 𝕜 g t y (fderiv_within 𝕜 f s x v) = fderiv_within 𝕜 (g ∘ f) s x v := -by { subst y, rw [fderiv_within.comp x hg hf h hxs], refl } - -/-- Ternary version of `fderiv_within.comp`, with equality assumptions of basepoints added, in - order to apply more easily as a rewrite from right-to-left. -/ -lemma fderiv_within.comp₃ {g' : G → G'} {g : F → G} {t : set F} {u : set G} {y : F} {y' : G} - (hg' : differentiable_within_at 𝕜 g' u y') (hg : differentiable_within_at 𝕜 g t y) - (hf : differentiable_within_at 𝕜 f s x) - (h2g : maps_to g t u) (h2f : maps_to f s t) - (h3g : g y = y') (h3f : f x = y) (hxs : unique_diff_within_at 𝕜 s x) : - fderiv_within 𝕜 (g' ∘ g ∘ f) s x = (fderiv_within 𝕜 g' u y').comp - ((fderiv_within 𝕜 g t y).comp (fderiv_within 𝕜 f s x)) := -begin - substs h3g h3f, - exact (hg'.has_fderiv_within_at.comp x - (hg.has_fderiv_within_at.comp x (hf.has_fderiv_within_at) h2f) $ h2g.comp h2f).fderiv_within hxs -end - -lemma fderiv.comp {g : F → G} - (hg : differentiable_at 𝕜 g (f x)) (hf : differentiable_at 𝕜 f x) : - fderiv 𝕜 (g ∘ f) x = (fderiv 𝕜 g (f x)).comp (fderiv 𝕜 f x) := -(hg.has_fderiv_at.comp x hf.has_fderiv_at).fderiv - -lemma fderiv.comp_fderiv_within {g : F → G} - (hg : differentiable_at 𝕜 g (f x)) (hf : differentiable_within_at 𝕜 f s x) - (hxs : unique_diff_within_at 𝕜 s x) : - fderiv_within 𝕜 (g ∘ f) s x = (fderiv 𝕜 g (f x)).comp (fderiv_within 𝕜 f s x) := -(hg.has_fderiv_at.comp_has_fderiv_within_at x hf.has_fderiv_within_at).fderiv_within hxs - -lemma differentiable_on.comp {g : F → G} {t : set F} - (hg : differentiable_on 𝕜 g t) (hf : differentiable_on 𝕜 f s) (st : maps_to f s t) : - differentiable_on 𝕜 (g ∘ f) s := -λx hx, differentiable_within_at.comp x (hg (f x) (st hx)) (hf x hx) st - -lemma differentiable.comp {g : F → G} (hg : differentiable 𝕜 g) (hf : differentiable 𝕜 f) : - differentiable 𝕜 (g ∘ f) := -λx, differentiable_at.comp x (hg (f x)) (hf x) - -lemma differentiable.comp_differentiable_on {g : F → G} (hg : differentiable 𝕜 g) - (hf : differentiable_on 𝕜 f s) : - differentiable_on 𝕜 (g ∘ f) s := -hg.differentiable_on.comp hf (maps_to_univ _ _) - -/-- The chain rule for derivatives in the sense of strict differentiability. -/ -protected lemma has_strict_fderiv_at.comp {g : F → G} {g' : F →L[𝕜] G} - (hg : has_strict_fderiv_at g g' (f x)) (hf : has_strict_fderiv_at f f' x) : - has_strict_fderiv_at (λ x, g (f x)) (g'.comp f') x := -((hg.comp_tendsto (hf.continuous_at.prod_map' hf.continuous_at)).trans_is_O hf.is_O_sub).triangle $ - by simpa only [g'.map_sub, f'.coe_comp'] using (g'.is_O_comp _ _).trans_is_o hf - -protected lemma differentiable.iterate {f : E → E} (hf : differentiable 𝕜 f) (n : ℕ) : - differentiable 𝕜 (f^[n]) := -nat.rec_on n differentiable_id (λ n ihn, ihn.comp hf) - -protected lemma differentiable_on.iterate {f : E → E} (hf : differentiable_on 𝕜 f s) - (hs : maps_to f s s) (n : ℕ) : - differentiable_on 𝕜 (f^[n]) s := -nat.rec_on n differentiable_on_id (λ n ihn, ihn.comp hf hs) - -variable {x} - -protected lemma has_fderiv_at_filter.iterate {f : E → E} {f' : E →L[𝕜] E} - (hf : has_fderiv_at_filter f f' x L) (hL : tendsto f L L) (hx : f x = x) (n : ℕ) : - has_fderiv_at_filter (f^[n]) (f'^n) x L := -begin - induction n with n ihn, - { exact has_fderiv_at_filter_id x L }, - { rw [function.iterate_succ, pow_succ'], - rw ← hx at ihn, - exact ihn.comp x hf hL } -end - -protected lemma has_fderiv_at.iterate {f : E → E} {f' : E →L[𝕜] E} - (hf : has_fderiv_at f f' x) (hx : f x = x) (n : ℕ) : - has_fderiv_at (f^[n]) (f'^n) x := -begin - refine hf.iterate _ hx n, - convert hf.continuous_at, - exact hx.symm -end - -protected lemma has_fderiv_within_at.iterate {f : E → E} {f' : E →L[𝕜] E} - (hf : has_fderiv_within_at f f' s x) (hx : f x = x) (hs : maps_to f s s) (n : ℕ) : - has_fderiv_within_at (f^[n]) (f'^n) s x := -begin - refine hf.iterate _ hx n, - convert tendsto_inf.2 ⟨hf.continuous_within_at, _⟩, - exacts [hx.symm, (tendsto_principal_principal.2 hs).mono_left inf_le_right] -end - -protected lemma has_strict_fderiv_at.iterate {f : E → E} {f' : E →L[𝕜] E} - (hf : has_strict_fderiv_at f f' x) (hx : f x = x) (n : ℕ) : - has_strict_fderiv_at (f^[n]) (f'^n) x := -begin - induction n with n ihn, - { exact has_strict_fderiv_at_id x }, - { rw [function.iterate_succ, pow_succ'], - rw ← hx at ihn, - exact ihn.comp x hf } end -protected lemma differentiable_at.iterate {f : E → E} (hf : differentiable_at 𝕜 f x) - (hx : f x = x) (n : ℕ) : - differentiable_at 𝕜 (f^[n]) x := -(hf.has_fderiv_at.iterate hx n).differentiable_at - -protected lemma differentiable_within_at.iterate {f : E → E} (hf : differentiable_within_at 𝕜 f s x) - (hx : f x = x) (hs : maps_to f s s) (n : ℕ) : - differentiable_within_at 𝕜 (f^[n]) s x := -(hf.has_fderiv_within_at.iterate hx hs n).differentiable_within_at - -end composition - -section cartesian_product -/-! ### Derivative of the cartesian product of two functions -/ - -section prod -variables {f₂ : E → G} {f₂' : E →L[𝕜] G} - -protected lemma has_strict_fderiv_at.prod - (hf₁ : has_strict_fderiv_at f₁ f₁' x) (hf₂ : has_strict_fderiv_at f₂ f₂' x) : - has_strict_fderiv_at (λx, (f₁ x, f₂ x)) (f₁'.prod f₂') x := -hf₁.prod_left hf₂ - -lemma has_fderiv_at_filter.prod - (hf₁ : has_fderiv_at_filter f₁ f₁' x L) (hf₂ : has_fderiv_at_filter f₂ f₂' x L) : - has_fderiv_at_filter (λx, (f₁ x, f₂ x)) (f₁'.prod f₂') x L := -hf₁.prod_left hf₂ - -lemma has_fderiv_within_at.prod - (hf₁ : has_fderiv_within_at f₁ f₁' s x) (hf₂ : has_fderiv_within_at f₂ f₂' s x) : - has_fderiv_within_at (λx, (f₁ x, f₂ x)) (f₁'.prod f₂') s x := -hf₁.prod hf₂ - -lemma has_fderiv_at.prod (hf₁ : has_fderiv_at f₁ f₁' x) (hf₂ : has_fderiv_at f₂ f₂' x) : - has_fderiv_at (λx, (f₁ x, f₂ x)) (f₁'.prod f₂') x := -hf₁.prod hf₂ - -lemma has_fderiv_at_prod_mk_left (e₀ : E) (f₀ : F) : - has_fderiv_at (λ e : E, (e, f₀)) (inl 𝕜 E F) e₀ := -(has_fderiv_at_id e₀).prod (has_fderiv_at_const f₀ e₀) - -lemma has_fderiv_at_prod_mk_right (e₀ : E) (f₀ : F) : - has_fderiv_at (λ f : F, (e₀, f)) (inr 𝕜 E F) f₀ := -(has_fderiv_at_const e₀ f₀).prod (has_fderiv_at_id f₀) - -lemma differentiable_within_at.prod - (hf₁ : differentiable_within_at 𝕜 f₁ s x) (hf₂ : differentiable_within_at 𝕜 f₂ s x) : - differentiable_within_at 𝕜 (λx:E, (f₁ x, f₂ x)) s x := -(hf₁.has_fderiv_within_at.prod hf₂.has_fderiv_within_at).differentiable_within_at - -@[simp] -lemma differentiable_at.prod (hf₁ : differentiable_at 𝕜 f₁ x) (hf₂ : differentiable_at 𝕜 f₂ x) : - differentiable_at 𝕜 (λx:E, (f₁ x, f₂ x)) x := -(hf₁.has_fderiv_at.prod hf₂.has_fderiv_at).differentiable_at - -lemma differentiable_on.prod (hf₁ : differentiable_on 𝕜 f₁ s) (hf₂ : differentiable_on 𝕜 f₂ s) : - differentiable_on 𝕜 (λx:E, (f₁ x, f₂ x)) s := -λx hx, differentiable_within_at.prod (hf₁ x hx) (hf₂ x hx) - -@[simp] -lemma differentiable.prod (hf₁ : differentiable 𝕜 f₁) (hf₂ : differentiable 𝕜 f₂) : - differentiable 𝕜 (λx:E, (f₁ x, f₂ x)) := -λ x, differentiable_at.prod (hf₁ x) (hf₂ x) - -lemma differentiable_at.fderiv_prod - (hf₁ : differentiable_at 𝕜 f₁ x) (hf₂ : differentiable_at 𝕜 f₂ x) : - fderiv 𝕜 (λx:E, (f₁ x, f₂ x)) x = (fderiv 𝕜 f₁ x).prod (fderiv 𝕜 f₂ x) := -(hf₁.has_fderiv_at.prod hf₂.has_fderiv_at).fderiv - -lemma differentiable_at.fderiv_within_prod - (hf₁ : differentiable_within_at 𝕜 f₁ s x) (hf₂ : differentiable_within_at 𝕜 f₂ s x) - (hxs : unique_diff_within_at 𝕜 s x) : - fderiv_within 𝕜 (λx:E, (f₁ x, f₂ x)) s x = - (fderiv_within 𝕜 f₁ s x).prod (fderiv_within 𝕜 f₂ s x) := -(hf₁.has_fderiv_within_at.prod hf₂.has_fderiv_within_at).fderiv_within hxs - -end prod - -section fst - -variables {f₂ : E → F × G} {f₂' : E →L[𝕜] F × G} {p : E × F} - -lemma has_strict_fderiv_at_fst : has_strict_fderiv_at (@prod.fst E F) (fst 𝕜 E F) p := -(fst 𝕜 E F).has_strict_fderiv_at - -protected lemma has_strict_fderiv_at.fst (h : has_strict_fderiv_at f₂ f₂' x) : - has_strict_fderiv_at (λ x, (f₂ x).1) ((fst 𝕜 F G).comp f₂') x := -has_strict_fderiv_at_fst.comp x h - -lemma has_fderiv_at_filter_fst {L : filter (E × F)} : - has_fderiv_at_filter (@prod.fst E F) (fst 𝕜 E F) p L := -(fst 𝕜 E F).has_fderiv_at_filter - -protected lemma has_fderiv_at_filter.fst (h : has_fderiv_at_filter f₂ f₂' x L) : - has_fderiv_at_filter (λ x, (f₂ x).1) ((fst 𝕜 F G).comp f₂') x L := -has_fderiv_at_filter_fst.comp x h tendsto_map - -lemma has_fderiv_at_fst : has_fderiv_at (@prod.fst E F) (fst 𝕜 E F) p := -has_fderiv_at_filter_fst - -protected lemma has_fderiv_at.fst (h : has_fderiv_at f₂ f₂' x) : - has_fderiv_at (λ x, (f₂ x).1) ((fst 𝕜 F G).comp f₂') x := -h.fst - -lemma has_fderiv_within_at_fst {s : set (E × F)} : - has_fderiv_within_at (@prod.fst E F) (fst 𝕜 E F) s p := -has_fderiv_at_filter_fst - -protected lemma has_fderiv_within_at.fst (h : has_fderiv_within_at f₂ f₂' s x) : - has_fderiv_within_at (λ x, (f₂ x).1) ((fst 𝕜 F G).comp f₂') s x := -h.fst - -lemma differentiable_at_fst : differentiable_at 𝕜 prod.fst p := -has_fderiv_at_fst.differentiable_at - -@[simp] protected lemma differentiable_at.fst (h : differentiable_at 𝕜 f₂ x) : - differentiable_at 𝕜 (λ x, (f₂ x).1) x := -differentiable_at_fst.comp x h - -lemma differentiable_fst : differentiable 𝕜 (prod.fst : E × F → E) := -λ x, differentiable_at_fst - -@[simp] protected lemma differentiable.fst (h : differentiable 𝕜 f₂) : - differentiable 𝕜 (λ x, (f₂ x).1) := -differentiable_fst.comp h - -lemma differentiable_within_at_fst {s : set (E × F)} : differentiable_within_at 𝕜 prod.fst s p := -differentiable_at_fst.differentiable_within_at - -protected lemma differentiable_within_at.fst (h : differentiable_within_at 𝕜 f₂ s x) : - differentiable_within_at 𝕜 (λ x, (f₂ x).1) s x := -differentiable_at_fst.comp_differentiable_within_at x h - -lemma differentiable_on_fst {s : set (E × F)} : differentiable_on 𝕜 prod.fst s := -differentiable_fst.differentiable_on - -protected lemma differentiable_on.fst (h : differentiable_on 𝕜 f₂ s) : - differentiable_on 𝕜 (λ x, (f₂ x).1) s := -differentiable_fst.comp_differentiable_on h - -lemma fderiv_fst : fderiv 𝕜 prod.fst p = fst 𝕜 E F := has_fderiv_at_fst.fderiv - -lemma fderiv.fst (h : differentiable_at 𝕜 f₂ x) : - fderiv 𝕜 (λ x, (f₂ x).1) x = (fst 𝕜 F G).comp (fderiv 𝕜 f₂ x) := -h.has_fderiv_at.fst.fderiv - -lemma fderiv_within_fst {s : set (E × F)} (hs : unique_diff_within_at 𝕜 s p) : - fderiv_within 𝕜 prod.fst s p = fst 𝕜 E F := -has_fderiv_within_at_fst.fderiv_within hs - -lemma fderiv_within.fst (hs : unique_diff_within_at 𝕜 s x) (h : differentiable_within_at 𝕜 f₂ s x) : - fderiv_within 𝕜 (λ x, (f₂ x).1) s x = (fst 𝕜 F G).comp (fderiv_within 𝕜 f₂ s x) := -h.has_fderiv_within_at.fst.fderiv_within hs - -end fst - -section snd - -variables {f₂ : E → F × G} {f₂' : E →L[𝕜] F × G} {p : E × F} - -lemma has_strict_fderiv_at_snd : has_strict_fderiv_at (@prod.snd E F) (snd 𝕜 E F) p := -(snd 𝕜 E F).has_strict_fderiv_at - -protected lemma has_strict_fderiv_at.snd (h : has_strict_fderiv_at f₂ f₂' x) : - has_strict_fderiv_at (λ x, (f₂ x).2) ((snd 𝕜 F G).comp f₂') x := -has_strict_fderiv_at_snd.comp x h - -lemma has_fderiv_at_filter_snd {L : filter (E × F)} : - has_fderiv_at_filter (@prod.snd E F) (snd 𝕜 E F) p L := -(snd 𝕜 E F).has_fderiv_at_filter - -protected lemma has_fderiv_at_filter.snd (h : has_fderiv_at_filter f₂ f₂' x L) : - has_fderiv_at_filter (λ x, (f₂ x).2) ((snd 𝕜 F G).comp f₂') x L := -has_fderiv_at_filter_snd.comp x h tendsto_map - -lemma has_fderiv_at_snd : has_fderiv_at (@prod.snd E F) (snd 𝕜 E F) p := -has_fderiv_at_filter_snd - -protected lemma has_fderiv_at.snd (h : has_fderiv_at f₂ f₂' x) : - has_fderiv_at (λ x, (f₂ x).2) ((snd 𝕜 F G).comp f₂') x := -h.snd - -lemma has_fderiv_within_at_snd {s : set (E × F)} : - has_fderiv_within_at (@prod.snd E F) (snd 𝕜 E F) s p := -has_fderiv_at_filter_snd - -protected lemma has_fderiv_within_at.snd (h : has_fderiv_within_at f₂ f₂' s x) : - has_fderiv_within_at (λ x, (f₂ x).2) ((snd 𝕜 F G).comp f₂') s x := -h.snd - -lemma differentiable_at_snd : differentiable_at 𝕜 prod.snd p := -has_fderiv_at_snd.differentiable_at - -@[simp] protected lemma differentiable_at.snd (h : differentiable_at 𝕜 f₂ x) : - differentiable_at 𝕜 (λ x, (f₂ x).2) x := -differentiable_at_snd.comp x h - -lemma differentiable_snd : differentiable 𝕜 (prod.snd : E × F → F) := -λ x, differentiable_at_snd - -@[simp] protected lemma differentiable.snd (h : differentiable 𝕜 f₂) : - differentiable 𝕜 (λ x, (f₂ x).2) := -differentiable_snd.comp h - -lemma differentiable_within_at_snd {s : set (E × F)} : differentiable_within_at 𝕜 prod.snd s p := -differentiable_at_snd.differentiable_within_at - -protected lemma differentiable_within_at.snd (h : differentiable_within_at 𝕜 f₂ s x) : - differentiable_within_at 𝕜 (λ x, (f₂ x).2) s x := -differentiable_at_snd.comp_differentiable_within_at x h - -lemma differentiable_on_snd {s : set (E × F)} : differentiable_on 𝕜 prod.snd s := -differentiable_snd.differentiable_on - -protected lemma differentiable_on.snd (h : differentiable_on 𝕜 f₂ s) : - differentiable_on 𝕜 (λ x, (f₂ x).2) s := -differentiable_snd.comp_differentiable_on h - -lemma fderiv_snd : fderiv 𝕜 prod.snd p = snd 𝕜 E F := has_fderiv_at_snd.fderiv - -lemma fderiv.snd (h : differentiable_at 𝕜 f₂ x) : - fderiv 𝕜 (λ x, (f₂ x).2) x = (snd 𝕜 F G).comp (fderiv 𝕜 f₂ x) := -h.has_fderiv_at.snd.fderiv - -lemma fderiv_within_snd {s : set (E × F)} (hs : unique_diff_within_at 𝕜 s p) : - fderiv_within 𝕜 prod.snd s p = snd 𝕜 E F := -has_fderiv_within_at_snd.fderiv_within hs - -lemma fderiv_within.snd (hs : unique_diff_within_at 𝕜 s x) (h : differentiable_within_at 𝕜 f₂ s x) : - fderiv_within 𝕜 (λ x, (f₂ x).2) s x = (snd 𝕜 F G).comp (fderiv_within 𝕜 f₂ s x) := -h.has_fderiv_within_at.snd.fderiv_within hs - -end snd - -section prod_map - -variables {f₂ : G → G'} {f₂' : G →L[𝕜] G'} {y : G} (p : E × G) - -protected theorem has_strict_fderiv_at.prod_map (hf : has_strict_fderiv_at f f' p.1) - (hf₂ : has_strict_fderiv_at f₂ f₂' p.2) : - has_strict_fderiv_at (prod.map f f₂) (f'.prod_map f₂') p := -(hf.comp p has_strict_fderiv_at_fst).prod (hf₂.comp p has_strict_fderiv_at_snd) - -protected theorem has_fderiv_at.prod_map (hf : has_fderiv_at f f' p.1) - (hf₂ : has_fderiv_at f₂ f₂' p.2) : - has_fderiv_at (prod.map f f₂) (f'.prod_map f₂') p := -(hf.comp p has_fderiv_at_fst).prod (hf₂.comp p has_fderiv_at_snd) - -@[simp] protected theorem differentiable_at.prod_map (hf : differentiable_at 𝕜 f p.1) - (hf₂ : differentiable_at 𝕜 f₂ p.2) : - differentiable_at 𝕜 (λ p : E × G, (f p.1, f₂ p.2)) p := -(hf.comp p differentiable_at_fst).prod (hf₂.comp p differentiable_at_snd) - -end prod_map - -end cartesian_product - -section const_smul - -variables {R : Type*} [semiring R] [module R F] [smul_comm_class 𝕜 R F] - [has_continuous_const_smul R F] - -/-! ### Derivative of a function multiplied by a constant -/ -theorem has_strict_fderiv_at.const_smul (h : has_strict_fderiv_at f f' x) (c : R) : - has_strict_fderiv_at (λ x, c • f x) (c • f') x := -(c • (1 : F →L[𝕜] F)).has_strict_fderiv_at.comp x h - -theorem has_fderiv_at_filter.const_smul (h : has_fderiv_at_filter f f' x L) (c : R) : - has_fderiv_at_filter (λ x, c • f x) (c • f') x L := -(c • (1 : F →L[𝕜] F)).has_fderiv_at_filter.comp x h tendsto_map - -theorem has_fderiv_within_at.const_smul (h : has_fderiv_within_at f f' s x) (c : R) : - has_fderiv_within_at (λ x, c • f x) (c • f') s x := -h.const_smul c - -theorem has_fderiv_at.const_smul (h : has_fderiv_at f f' x) (c : R) : - has_fderiv_at (λ x, c • f x) (c • f') x := -h.const_smul c - -lemma differentiable_within_at.const_smul (h : differentiable_within_at 𝕜 f s x) (c : R) : - differentiable_within_at 𝕜 (λy, c • f y) s x := -(h.has_fderiv_within_at.const_smul c).differentiable_within_at - -lemma differentiable_at.const_smul (h : differentiable_at 𝕜 f x) (c : R) : - differentiable_at 𝕜 (λy, c • f y) x := -(h.has_fderiv_at.const_smul c).differentiable_at - -lemma differentiable_on.const_smul (h : differentiable_on 𝕜 f s) (c : R) : - differentiable_on 𝕜 (λy, c • f y) s := -λx hx, (h x hx).const_smul c - -lemma differentiable.const_smul (h : differentiable 𝕜 f) (c : R) : - differentiable 𝕜 (λy, c • f y) := -λx, (h x).const_smul c - -lemma fderiv_within_const_smul (hxs : unique_diff_within_at 𝕜 s x) - (h : differentiable_within_at 𝕜 f s x) (c : R) : - fderiv_within 𝕜 (λy, c • f y) s x = c • fderiv_within 𝕜 f s x := -(h.has_fderiv_within_at.const_smul c).fderiv_within hxs - -lemma fderiv_const_smul (h : differentiable_at 𝕜 f x) (c : R) : - fderiv 𝕜 (λy, c • f y) x = c • fderiv 𝕜 f x := -(h.has_fderiv_at.const_smul c).fderiv - -end const_smul - -section add - -/-! ### Derivative of the sum of two functions -/ - -theorem has_strict_fderiv_at.add (hf : has_strict_fderiv_at f f' x) - (hg : has_strict_fderiv_at g g' x) : - has_strict_fderiv_at (λ y, f y + g y) (f' + g') x := -(hf.add hg).congr_left $ λ y, - by { simp only [linear_map.sub_apply, linear_map.add_apply, map_sub, map_add, add_apply], abel } - -theorem has_fderiv_at_filter.add - (hf : has_fderiv_at_filter f f' x L) (hg : has_fderiv_at_filter g g' x L) : - has_fderiv_at_filter (λ y, f y + g y) (f' + g') x L := -(hf.add hg).congr_left $ λ _, - by { simp only [linear_map.sub_apply, linear_map.add_apply, map_sub, map_add, add_apply], abel } - -theorem has_fderiv_within_at.add - (hf : has_fderiv_within_at f f' s x) (hg : has_fderiv_within_at g g' s x) : - has_fderiv_within_at (λ y, f y + g y) (f' + g') s x := -hf.add hg - -theorem has_fderiv_at.add - (hf : has_fderiv_at f f' x) (hg : has_fderiv_at g g' x) : - has_fderiv_at (λ x, f x + g x) (f' + g') x := -hf.add hg - -lemma differentiable_within_at.add - (hf : differentiable_within_at 𝕜 f s x) (hg : differentiable_within_at 𝕜 g s x) : - differentiable_within_at 𝕜 (λ y, f y + g y) s x := -(hf.has_fderiv_within_at.add hg.has_fderiv_within_at).differentiable_within_at - -@[simp] lemma differentiable_at.add - (hf : differentiable_at 𝕜 f x) (hg : differentiable_at 𝕜 g x) : - differentiable_at 𝕜 (λ y, f y + g y) x := -(hf.has_fderiv_at.add hg.has_fderiv_at).differentiable_at - -lemma differentiable_on.add - (hf : differentiable_on 𝕜 f s) (hg : differentiable_on 𝕜 g s) : - differentiable_on 𝕜 (λy, f y + g y) s := -λx hx, (hf x hx).add (hg x hx) - -@[simp] lemma differentiable.add - (hf : differentiable 𝕜 f) (hg : differentiable 𝕜 g) : - differentiable 𝕜 (λy, f y + g y) := -λx, (hf x).add (hg x) - -lemma fderiv_within_add (hxs : unique_diff_within_at 𝕜 s x) - (hf : differentiable_within_at 𝕜 f s x) (hg : differentiable_within_at 𝕜 g s x) : - fderiv_within 𝕜 (λy, f y + g y) s x = fderiv_within 𝕜 f s x + fderiv_within 𝕜 g s x := -(hf.has_fderiv_within_at.add hg.has_fderiv_within_at).fderiv_within hxs - -lemma fderiv_add - (hf : differentiable_at 𝕜 f x) (hg : differentiable_at 𝕜 g x) : - fderiv 𝕜 (λy, f y + g y) x = fderiv 𝕜 f x + fderiv 𝕜 g x := -(hf.has_fderiv_at.add hg.has_fderiv_at).fderiv - -theorem has_strict_fderiv_at.add_const (hf : has_strict_fderiv_at f f' x) (c : F) : - has_strict_fderiv_at (λ y, f y + c) f' x := -add_zero f' ▸ hf.add (has_strict_fderiv_at_const _ _) - -theorem has_fderiv_at_filter.add_const - (hf : has_fderiv_at_filter f f' x L) (c : F) : - has_fderiv_at_filter (λ y, f y + c) f' x L := -add_zero f' ▸ hf.add (has_fderiv_at_filter_const _ _ _) - -theorem has_fderiv_within_at.add_const - (hf : has_fderiv_within_at f f' s x) (c : F) : - has_fderiv_within_at (λ y, f y + c) f' s x := -hf.add_const c - -theorem has_fderiv_at.add_const (hf : has_fderiv_at f f' x) (c : F): - has_fderiv_at (λ x, f x + c) f' x := -hf.add_const c - -lemma differentiable_within_at.add_const - (hf : differentiable_within_at 𝕜 f s x) (c : F) : - differentiable_within_at 𝕜 (λ y, f y + c) s x := -(hf.has_fderiv_within_at.add_const c).differentiable_within_at - -@[simp] lemma differentiable_within_at_add_const_iff (c : F) : - differentiable_within_at 𝕜 (λ y, f y + c) s x ↔ differentiable_within_at 𝕜 f s x := -⟨λ h, by simpa using h.add_const (-c), λ h, h.add_const c⟩ - -lemma differentiable_at.add_const - (hf : differentiable_at 𝕜 f x) (c : F) : - differentiable_at 𝕜 (λ y, f y + c) x := -(hf.has_fderiv_at.add_const c).differentiable_at - -@[simp] lemma differentiable_at_add_const_iff (c : F) : - differentiable_at 𝕜 (λ y, f y + c) x ↔ differentiable_at 𝕜 f x := -⟨λ h, by simpa using h.add_const (-c), λ h, h.add_const c⟩ - -lemma differentiable_on.add_const - (hf : differentiable_on 𝕜 f s) (c : F) : - differentiable_on 𝕜 (λy, f y + c) s := -λx hx, (hf x hx).add_const c - -@[simp] lemma differentiable_on_add_const_iff (c : F) : - differentiable_on 𝕜 (λ y, f y + c) s ↔ differentiable_on 𝕜 f s := -⟨λ h, by simpa using h.add_const (-c), λ h, h.add_const c⟩ - -lemma differentiable.add_const - (hf : differentiable 𝕜 f) (c : F) : - differentiable 𝕜 (λy, f y + c) := -λx, (hf x).add_const c - -@[simp] lemma differentiable_add_const_iff (c : F) : - differentiable 𝕜 (λ y, f y + c) ↔ differentiable 𝕜 f := -⟨λ h, by simpa using h.add_const (-c), λ h, h.add_const c⟩ - -lemma fderiv_within_add_const (hxs : unique_diff_within_at 𝕜 s x) (c : F) : - fderiv_within 𝕜 (λy, f y + c) s x = fderiv_within 𝕜 f s x := -if hf : differentiable_within_at 𝕜 f s x -then (hf.has_fderiv_within_at.add_const c).fderiv_within hxs -else by { rw [fderiv_within_zero_of_not_differentiable_within_at hf, - fderiv_within_zero_of_not_differentiable_within_at], simpa } - -lemma fderiv_add_const (c : F) : fderiv 𝕜 (λy, f y + c) x = fderiv 𝕜 f x := -by simp only [← fderiv_within_univ, fderiv_within_add_const unique_diff_within_at_univ] - -theorem has_strict_fderiv_at.const_add (hf : has_strict_fderiv_at f f' x) (c : F) : - has_strict_fderiv_at (λ y, c + f y) f' x := -zero_add f' ▸ (has_strict_fderiv_at_const _ _).add hf - -theorem has_fderiv_at_filter.const_add - (hf : has_fderiv_at_filter f f' x L) (c : F) : - has_fderiv_at_filter (λ y, c + f y) f' x L := -zero_add f' ▸ (has_fderiv_at_filter_const _ _ _).add hf - -theorem has_fderiv_within_at.const_add - (hf : has_fderiv_within_at f f' s x) (c : F) : - has_fderiv_within_at (λ y, c + f y) f' s x := -hf.const_add c - -theorem has_fderiv_at.const_add - (hf : has_fderiv_at f f' x) (c : F): - has_fderiv_at (λ x, c + f x) f' x := -hf.const_add c - -lemma differentiable_within_at.const_add - (hf : differentiable_within_at 𝕜 f s x) (c : F) : - differentiable_within_at 𝕜 (λ y, c + f y) s x := -(hf.has_fderiv_within_at.const_add c).differentiable_within_at - -@[simp] lemma differentiable_within_at_const_add_iff (c : F) : - differentiable_within_at 𝕜 (λ y, c + f y) s x ↔ differentiable_within_at 𝕜 f s x := -⟨λ h, by simpa using h.const_add (-c), λ h, h.const_add c⟩ - -lemma differentiable_at.const_add - (hf : differentiable_at 𝕜 f x) (c : F) : - differentiable_at 𝕜 (λ y, c + f y) x := -(hf.has_fderiv_at.const_add c).differentiable_at - -@[simp] lemma differentiable_at_const_add_iff (c : F) : - differentiable_at 𝕜 (λ y, c + f y) x ↔ differentiable_at 𝕜 f x := -⟨λ h, by simpa using h.const_add (-c), λ h, h.const_add c⟩ - -lemma differentiable_on.const_add (hf : differentiable_on 𝕜 f s) (c : F) : - differentiable_on 𝕜 (λy, c + f y) s := -λx hx, (hf x hx).const_add c - -@[simp] lemma differentiable_on_const_add_iff (c : F) : - differentiable_on 𝕜 (λ y, c + f y) s ↔ differentiable_on 𝕜 f s := -⟨λ h, by simpa using h.const_add (-c), λ h, h.const_add c⟩ - -lemma differentiable.const_add (hf : differentiable 𝕜 f) (c : F) : - differentiable 𝕜 (λy, c + f y) := -λx, (hf x).const_add c - -@[simp] lemma differentiable_const_add_iff (c : F) : - differentiable 𝕜 (λ y, c + f y) ↔ differentiable 𝕜 f := -⟨λ h, by simpa using h.const_add (-c), λ h, h.const_add c⟩ - -lemma fderiv_within_const_add (hxs : unique_diff_within_at 𝕜 s x) (c : F) : - fderiv_within 𝕜 (λy, c + f y) s x = fderiv_within 𝕜 f s x := -by simpa only [add_comm] using fderiv_within_add_const hxs c - -lemma fderiv_const_add (c : F) : fderiv 𝕜 (λy, c + f y) x = fderiv 𝕜 f x := -by simp only [add_comm c, fderiv_add_const] - -end add - -section sum -/-! ### Derivative of a finite sum of functions -/ - -open_locale big_operators - -variables {ι : Type*} {u : finset ι} {A : ι → (E → F)} {A' : ι → (E →L[𝕜] F)} - -theorem has_strict_fderiv_at.sum (h : ∀ i ∈ u, has_strict_fderiv_at (A i) (A' i) x) : - has_strict_fderiv_at (λ y, ∑ i in u, A i y) (∑ i in u, A' i) x := -begin - dsimp [has_strict_fderiv_at] at *, - convert is_o.sum h, - simp [finset.sum_sub_distrib, continuous_linear_map.sum_apply] -end - -theorem has_fderiv_at_filter.sum (h : ∀ i ∈ u, has_fderiv_at_filter (A i) (A' i) x L) : - has_fderiv_at_filter (λ y, ∑ i in u, A i y) (∑ i in u, A' i) x L := -begin - dsimp [has_fderiv_at_filter] at *, - convert is_o.sum h, - simp [continuous_linear_map.sum_apply] -end - -theorem has_fderiv_within_at.sum (h : ∀ i ∈ u, has_fderiv_within_at (A i) (A' i) s x) : - has_fderiv_within_at (λ y, ∑ i in u, A i y) (∑ i in u, A' i) s x := -has_fderiv_at_filter.sum h - -theorem has_fderiv_at.sum (h : ∀ i ∈ u, has_fderiv_at (A i) (A' i) x) : - has_fderiv_at (λ y, ∑ i in u, A i y) (∑ i in u, A' i) x := -has_fderiv_at_filter.sum h - -theorem differentiable_within_at.sum (h : ∀ i ∈ u, differentiable_within_at 𝕜 (A i) s x) : - differentiable_within_at 𝕜 (λ y, ∑ i in u, A i y) s x := -has_fderiv_within_at.differentiable_within_at $ has_fderiv_within_at.sum $ -λ i hi, (h i hi).has_fderiv_within_at - -@[simp] theorem differentiable_at.sum (h : ∀ i ∈ u, differentiable_at 𝕜 (A i) x) : - differentiable_at 𝕜 (λ y, ∑ i in u, A i y) x := -has_fderiv_at.differentiable_at $ has_fderiv_at.sum $ λ i hi, (h i hi).has_fderiv_at - -theorem differentiable_on.sum (h : ∀ i ∈ u, differentiable_on 𝕜 (A i) s) : - differentiable_on 𝕜 (λ y, ∑ i in u, A i y) s := -λ x hx, differentiable_within_at.sum $ λ i hi, h i hi x hx - -@[simp] theorem differentiable.sum (h : ∀ i ∈ u, differentiable 𝕜 (A i)) : - differentiable 𝕜 (λ y, ∑ i in u, A i y) := -λ x, differentiable_at.sum $ λ i hi, h i hi x - -theorem fderiv_within_sum (hxs : unique_diff_within_at 𝕜 s x) - (h : ∀ i ∈ u, differentiable_within_at 𝕜 (A i) s x) : - fderiv_within 𝕜 (λ y, ∑ i in u, A i y) s x = (∑ i in u, fderiv_within 𝕜 (A i) s x) := -(has_fderiv_within_at.sum (λ i hi, (h i hi).has_fderiv_within_at)).fderiv_within hxs - -theorem fderiv_sum (h : ∀ i ∈ u, differentiable_at 𝕜 (A i) x) : - fderiv 𝕜 (λ y, ∑ i in u, A i y) x = (∑ i in u, fderiv 𝕜 (A i) x) := -(has_fderiv_at.sum (λ i hi, (h i hi).has_fderiv_at)).fderiv - -end sum - -section pi - -/-! -### Derivatives of functions `f : E → Π i, F' i` - -In this section we formulate `has_*fderiv*_pi` theorems as `iff`s, and provide two versions of each -theorem: - -* the version without `'` deals with `φ : Π i, E → F' i` and `φ' : Π i, E →L[𝕜] F' i` - and is designed to deduce differentiability of `λ x i, φ i x` from differentiability - of each `φ i`; -* the version with `'` deals with `Φ : E → Π i, F' i` and `Φ' : E →L[𝕜] Π i, F' i` - and is designed to deduce differentiability of the components `λ x, Φ x i` from - differentiability of `Φ`. --/ - -variables {ι : Type*} [fintype ι] {F' : ι → Type*} [Π i, normed_add_comm_group (F' i)] - [Π i, normed_space 𝕜 (F' i)] {φ : Π i, E → F' i} {φ' : Π i, E →L[𝕜] F' i} - {Φ : E → Π i, F' i} {Φ' : E →L[𝕜] Π i, F' i} - -@[simp] lemma has_strict_fderiv_at_pi' : - has_strict_fderiv_at Φ Φ' x ↔ - ∀ i, has_strict_fderiv_at (λ x, Φ x i) ((proj i).comp Φ') x := -begin - simp only [has_strict_fderiv_at, continuous_linear_map.coe_pi], - exact is_o_pi -end - -@[simp] lemma has_strict_fderiv_at_pi : - has_strict_fderiv_at (λ x i, φ i x) (continuous_linear_map.pi φ') x ↔ - ∀ i, has_strict_fderiv_at (φ i) (φ' i) x := -has_strict_fderiv_at_pi' - -@[simp] lemma has_fderiv_at_filter_pi' : - has_fderiv_at_filter Φ Φ' x L ↔ - ∀ i, has_fderiv_at_filter (λ x, Φ x i) ((proj i).comp Φ') x L := -begin - simp only [has_fderiv_at_filter, continuous_linear_map.coe_pi], - exact is_o_pi -end - -lemma has_fderiv_at_filter_pi : - has_fderiv_at_filter (λ x i, φ i x) (continuous_linear_map.pi φ') x L ↔ - ∀ i, has_fderiv_at_filter (φ i) (φ' i) x L := -has_fderiv_at_filter_pi' - -@[simp] lemma has_fderiv_at_pi' : - has_fderiv_at Φ Φ' x ↔ - ∀ i, has_fderiv_at (λ x, Φ x i) ((proj i).comp Φ') x := -has_fderiv_at_filter_pi' - -lemma has_fderiv_at_pi : - has_fderiv_at (λ x i, φ i x) (continuous_linear_map.pi φ') x ↔ - ∀ i, has_fderiv_at (φ i) (φ' i) x := -has_fderiv_at_filter_pi - -@[simp] lemma has_fderiv_within_at_pi' : - has_fderiv_within_at Φ Φ' s x ↔ - ∀ i, has_fderiv_within_at (λ x, Φ x i) ((proj i).comp Φ') s x := -has_fderiv_at_filter_pi' - -lemma has_fderiv_within_at_pi : - has_fderiv_within_at (λ x i, φ i x) (continuous_linear_map.pi φ') s x ↔ - ∀ i, has_fderiv_within_at (φ i) (φ' i) s x := -has_fderiv_at_filter_pi - -@[simp] lemma differentiable_within_at_pi : - differentiable_within_at 𝕜 Φ s x ↔ - ∀ i, differentiable_within_at 𝕜 (λ x, Φ x i) s x := -⟨λ h i, (has_fderiv_within_at_pi'.1 h.has_fderiv_within_at i).differentiable_within_at, - λ h, (has_fderiv_within_at_pi.2 (λ i, (h i).has_fderiv_within_at)).differentiable_within_at⟩ - -@[simp] lemma differentiable_at_pi : - differentiable_at 𝕜 Φ x ↔ ∀ i, differentiable_at 𝕜 (λ x, Φ x i) x := -⟨λ h i, (has_fderiv_at_pi'.1 h.has_fderiv_at i).differentiable_at, - λ h, (has_fderiv_at_pi.2 (λ i, (h i).has_fderiv_at)).differentiable_at⟩ - -lemma differentiable_on_pi : - differentiable_on 𝕜 Φ s ↔ ∀ i, differentiable_on 𝕜 (λ x, Φ x i) s := -⟨λ h i x hx, differentiable_within_at_pi.1 (h x hx) i, - λ h x hx, differentiable_within_at_pi.2 (λ i, h i x hx)⟩ - -lemma differentiable_pi : - differentiable 𝕜 Φ ↔ ∀ i, differentiable 𝕜 (λ x, Φ x i) := -⟨λ h i x, differentiable_at_pi.1 (h x) i, λ h x, differentiable_at_pi.2 (λ i, h i x)⟩ - --- TODO: find out which version (`φ` or `Φ`) works better with `rw`/`simp` -lemma fderiv_within_pi (h : ∀ i, differentiable_within_at 𝕜 (φ i) s x) - (hs : unique_diff_within_at 𝕜 s x) : - fderiv_within 𝕜 (λ x i, φ i x) s x = pi (λ i, fderiv_within 𝕜 (φ i) s x) := -(has_fderiv_within_at_pi.2 (λ i, (h i).has_fderiv_within_at)).fderiv_within hs - -lemma fderiv_pi (h : ∀ i, differentiable_at 𝕜 (φ i) x) : - fderiv 𝕜 (λ x i, φ i x) x = pi (λ i, fderiv 𝕜 (φ i) x) := -(has_fderiv_at_pi.2 (λ i, (h i).has_fderiv_at)).fderiv - -end pi - -section neg -/-! ### Derivative of the negative of a function -/ - -theorem has_strict_fderiv_at.neg (h : has_strict_fderiv_at f f' x) : - has_strict_fderiv_at (λ x, -f x) (-f') x := -(-1 : F →L[𝕜] F).has_strict_fderiv_at.comp x h - -theorem has_fderiv_at_filter.neg (h : has_fderiv_at_filter f f' x L) : - has_fderiv_at_filter (λ x, -f x) (-f') x L := -(-1 : F →L[𝕜] F).has_fderiv_at_filter.comp x h tendsto_map - -theorem has_fderiv_within_at.neg (h : has_fderiv_within_at f f' s x) : - has_fderiv_within_at (λ x, -f x) (-f') s x := -h.neg - -theorem has_fderiv_at.neg (h : has_fderiv_at f f' x) : - has_fderiv_at (λ x, -f x) (-f') x := -h.neg - -lemma differentiable_within_at.neg (h : differentiable_within_at 𝕜 f s x) : - differentiable_within_at 𝕜 (λy, -f y) s x := -h.has_fderiv_within_at.neg.differentiable_within_at - -@[simp] lemma differentiable_within_at_neg_iff : - differentiable_within_at 𝕜 (λy, -f y) s x ↔ differentiable_within_at 𝕜 f s x := -⟨λ h, by simpa only [neg_neg] using h.neg, λ h, h.neg⟩ - -lemma differentiable_at.neg (h : differentiable_at 𝕜 f x) : - differentiable_at 𝕜 (λy, -f y) x := -h.has_fderiv_at.neg.differentiable_at - -@[simp] lemma differentiable_at_neg_iff : - differentiable_at 𝕜 (λy, -f y) x ↔ differentiable_at 𝕜 f x := -⟨λ h, by simpa only [neg_neg] using h.neg, λ h, h.neg⟩ - -lemma differentiable_on.neg (h : differentiable_on 𝕜 f s) : - differentiable_on 𝕜 (λy, -f y) s := -λx hx, (h x hx).neg - -@[simp] lemma differentiable_on_neg_iff : - differentiable_on 𝕜 (λy, -f y) s ↔ differentiable_on 𝕜 f s := -⟨λ h, by simpa only [neg_neg] using h.neg, λ h, h.neg⟩ - -lemma differentiable.neg (h : differentiable 𝕜 f) : - differentiable 𝕜 (λy, -f y) := -λx, (h x).neg - -@[simp] lemma differentiable_neg_iff : differentiable 𝕜 (λy, -f y) ↔ differentiable 𝕜 f := -⟨λ h, by simpa only [neg_neg] using h.neg, λ h, h.neg⟩ - -lemma fderiv_within_neg (hxs : unique_diff_within_at 𝕜 s x) : - fderiv_within 𝕜 (λy, -f y) s x = - fderiv_within 𝕜 f s x := -if h : differentiable_within_at 𝕜 f s x -then h.has_fderiv_within_at.neg.fderiv_within hxs -else by { rw [fderiv_within_zero_of_not_differentiable_within_at h, - fderiv_within_zero_of_not_differentiable_within_at, neg_zero], simpa } - -@[simp] lemma fderiv_neg : fderiv 𝕜 (λy, -f y) x = - fderiv 𝕜 f x := -by simp only [← fderiv_within_univ, fderiv_within_neg unique_diff_within_at_univ] - -end neg - -section sub -/-! ### Derivative of the difference of two functions -/ - -theorem has_strict_fderiv_at.sub - (hf : has_strict_fderiv_at f f' x) (hg : has_strict_fderiv_at g g' x) : - has_strict_fderiv_at (λ x, f x - g x) (f' - g') x := -by simpa only [sub_eq_add_neg] using hf.add hg.neg - -theorem has_fderiv_at_filter.sub - (hf : has_fderiv_at_filter f f' x L) (hg : has_fderiv_at_filter g g' x L) : - has_fderiv_at_filter (λ x, f x - g x) (f' - g') x L := -by simpa only [sub_eq_add_neg] using hf.add hg.neg - -theorem has_fderiv_within_at.sub - (hf : has_fderiv_within_at f f' s x) (hg : has_fderiv_within_at g g' s x) : - has_fderiv_within_at (λ x, f x - g x) (f' - g') s x := -hf.sub hg - -theorem has_fderiv_at.sub - (hf : has_fderiv_at f f' x) (hg : has_fderiv_at g g' x) : - has_fderiv_at (λ x, f x - g x) (f' - g') x := -hf.sub hg - -lemma differentiable_within_at.sub - (hf : differentiable_within_at 𝕜 f s x) (hg : differentiable_within_at 𝕜 g s x) : - differentiable_within_at 𝕜 (λ y, f y - g y) s x := -(hf.has_fderiv_within_at.sub hg.has_fderiv_within_at).differentiable_within_at - -@[simp] lemma differentiable_at.sub - (hf : differentiable_at 𝕜 f x) (hg : differentiable_at 𝕜 g x) : - differentiable_at 𝕜 (λ y, f y - g y) x := -(hf.has_fderiv_at.sub hg.has_fderiv_at).differentiable_at - -lemma differentiable_on.sub - (hf : differentiable_on 𝕜 f s) (hg : differentiable_on 𝕜 g s) : - differentiable_on 𝕜 (λy, f y - g y) s := -λx hx, (hf x hx).sub (hg x hx) - -@[simp] lemma differentiable.sub - (hf : differentiable 𝕜 f) (hg : differentiable 𝕜 g) : - differentiable 𝕜 (λy, f y - g y) := -λx, (hf x).sub (hg x) - -lemma fderiv_within_sub (hxs : unique_diff_within_at 𝕜 s x) - (hf : differentiable_within_at 𝕜 f s x) (hg : differentiable_within_at 𝕜 g s x) : - fderiv_within 𝕜 (λy, f y - g y) s x = fderiv_within 𝕜 f s x - fderiv_within 𝕜 g s x := -(hf.has_fderiv_within_at.sub hg.has_fderiv_within_at).fderiv_within hxs - -lemma fderiv_sub - (hf : differentiable_at 𝕜 f x) (hg : differentiable_at 𝕜 g x) : - fderiv 𝕜 (λy, f y - g y) x = fderiv 𝕜 f x - fderiv 𝕜 g x := -(hf.has_fderiv_at.sub hg.has_fderiv_at).fderiv - -theorem has_strict_fderiv_at.sub_const - (hf : has_strict_fderiv_at f f' x) (c : F) : - has_strict_fderiv_at (λ x, f x - c) f' x := -by simpa only [sub_eq_add_neg] using hf.add_const (-c) - -theorem has_fderiv_at_filter.sub_const - (hf : has_fderiv_at_filter f f' x L) (c : F) : - has_fderiv_at_filter (λ x, f x - c) f' x L := -by simpa only [sub_eq_add_neg] using hf.add_const (-c) - -theorem has_fderiv_within_at.sub_const - (hf : has_fderiv_within_at f f' s x) (c : F) : - has_fderiv_within_at (λ x, f x - c) f' s x := -hf.sub_const c - -theorem has_fderiv_at.sub_const - (hf : has_fderiv_at f f' x) (c : F) : - has_fderiv_at (λ x, f x - c) f' x := -hf.sub_const c - -lemma differentiable_within_at.sub_const - (hf : differentiable_within_at 𝕜 f s x) (c : F) : - differentiable_within_at 𝕜 (λ y, f y - c) s x := -(hf.has_fderiv_within_at.sub_const c).differentiable_within_at - -@[simp] lemma differentiable_within_at_sub_const_iff (c : F) : - differentiable_within_at 𝕜 (λ y, f y - c) s x ↔ differentiable_within_at 𝕜 f s x := -by simp only [sub_eq_add_neg, differentiable_within_at_add_const_iff] - -lemma differentiable_at.sub_const (hf : differentiable_at 𝕜 f x) (c : F) : - differentiable_at 𝕜 (λ y, f y - c) x := -(hf.has_fderiv_at.sub_const c).differentiable_at - -@[simp] lemma differentiable_at_sub_const_iff (c : F) : - differentiable_at 𝕜 (λ y, f y - c) x ↔ differentiable_at 𝕜 f x := -by simp only [sub_eq_add_neg, differentiable_at_add_const_iff] - -lemma differentiable_on.sub_const (hf : differentiable_on 𝕜 f s) (c : F) : - differentiable_on 𝕜 (λy, f y - c) s := -λx hx, (hf x hx).sub_const c - -@[simp] lemma differentiable_on_sub_const_iff (c : F) : - differentiable_on 𝕜 (λ y, f y - c) s ↔ differentiable_on 𝕜 f s := -by simp only [sub_eq_add_neg, differentiable_on_add_const_iff] - -lemma differentiable.sub_const (hf : differentiable 𝕜 f) (c : F) : - differentiable 𝕜 (λy, f y - c) := -λx, (hf x).sub_const c - -@[simp] lemma differentiable_sub_const_iff (c : F) : - differentiable 𝕜 (λ y, f y - c) ↔ differentiable 𝕜 f := -by simp only [sub_eq_add_neg, differentiable_add_const_iff] - -lemma fderiv_within_sub_const (hxs : unique_diff_within_at 𝕜 s x) (c : F) : - fderiv_within 𝕜 (λy, f y - c) s x = fderiv_within 𝕜 f s x := -by simp only [sub_eq_add_neg, fderiv_within_add_const hxs] - -lemma fderiv_sub_const (c : F) : fderiv 𝕜 (λy, f y - c) x = fderiv 𝕜 f x := -by simp only [sub_eq_add_neg, fderiv_add_const] - -theorem has_strict_fderiv_at.const_sub - (hf : has_strict_fderiv_at f f' x) (c : F) : - has_strict_fderiv_at (λ x, c - f x) (-f') x := -by simpa only [sub_eq_add_neg] using hf.neg.const_add c - -theorem has_fderiv_at_filter.const_sub - (hf : has_fderiv_at_filter f f' x L) (c : F) : - has_fderiv_at_filter (λ x, c - f x) (-f') x L := -by simpa only [sub_eq_add_neg] using hf.neg.const_add c - -theorem has_fderiv_within_at.const_sub - (hf : has_fderiv_within_at f f' s x) (c : F) : - has_fderiv_within_at (λ x, c - f x) (-f') s x := -hf.const_sub c - -theorem has_fderiv_at.const_sub - (hf : has_fderiv_at f f' x) (c : F) : - has_fderiv_at (λ x, c - f x) (-f') x := -hf.const_sub c - -lemma differentiable_within_at.const_sub - (hf : differentiable_within_at 𝕜 f s x) (c : F) : - differentiable_within_at 𝕜 (λ y, c - f y) s x := -(hf.has_fderiv_within_at.const_sub c).differentiable_within_at - -@[simp] lemma differentiable_within_at_const_sub_iff (c : F) : - differentiable_within_at 𝕜 (λ y, c - f y) s x ↔ differentiable_within_at 𝕜 f s x := -by simp [sub_eq_add_neg] - -lemma differentiable_at.const_sub - (hf : differentiable_at 𝕜 f x) (c : F) : - differentiable_at 𝕜 (λ y, c - f y) x := -(hf.has_fderiv_at.const_sub c).differentiable_at - -@[simp] lemma differentiable_at_const_sub_iff (c : F) : - differentiable_at 𝕜 (λ y, c - f y) x ↔ differentiable_at 𝕜 f x := -by simp [sub_eq_add_neg] - -lemma differentiable_on.const_sub (hf : differentiable_on 𝕜 f s) (c : F) : - differentiable_on 𝕜 (λy, c - f y) s := -λx hx, (hf x hx).const_sub c - -@[simp] lemma differentiable_on_const_sub_iff (c : F) : - differentiable_on 𝕜 (λ y, c - f y) s ↔ differentiable_on 𝕜 f s := -by simp [sub_eq_add_neg] - -lemma differentiable.const_sub (hf : differentiable 𝕜 f) (c : F) : - differentiable 𝕜 (λy, c - f y) := -λx, (hf x).const_sub c - -@[simp] lemma differentiable_const_sub_iff (c : F) : - differentiable 𝕜 (λ y, c - f y) ↔ differentiable 𝕜 f := -by simp [sub_eq_add_neg] - -lemma fderiv_within_const_sub (hxs : unique_diff_within_at 𝕜 s x) (c : F) : - fderiv_within 𝕜 (λy, c - f y) s x = -fderiv_within 𝕜 f s x := -by simp only [sub_eq_add_neg, fderiv_within_const_add, fderiv_within_neg, hxs] - -lemma fderiv_const_sub (c : F) : fderiv 𝕜 (λy, c - f y) x = -fderiv 𝕜 f x := -by simp only [← fderiv_within_univ, fderiv_within_const_sub unique_diff_within_at_univ] - -end sub - -section bilinear_map -/-! ### Derivative of a bounded bilinear map -/ - -variables {b : E × F → G} {u : set (E × F)} - -open normed_field - -lemma is_bounded_bilinear_map.has_strict_fderiv_at (h : is_bounded_bilinear_map 𝕜 b) (p : E × F) : - has_strict_fderiv_at b (h.deriv p) p := -begin - rw has_strict_fderiv_at, - set T := (E × F) × (E × F), - have : (λ q : T, b (q.1 - q.2)) =o[𝓝 (p, p)] (λ q : T, ‖q.1 - q.2‖ * 1), - { refine (h.is_O'.comp_tendsto le_top).trans_is_o _, - simp only [(∘)], - refine (is_O_refl (λ q : T, ‖q.1 - q.2‖) _).mul_is_o (is_o.norm_left $ (is_o_one_iff _).2 _), - rw [← sub_self p], - exact continuous_at_fst.sub continuous_at_snd }, - simp only [mul_one, is_o_norm_right] at this, - refine (is_o.congr_of_sub _).1 this, clear this, - convert_to (λ q : T, h.deriv (p - q.2) (q.1 - q.2)) =o[𝓝 (p, p)] (λ q : T, q.1 - q.2), - { ext ⟨⟨x₁, y₁⟩, ⟨x₂, y₂⟩⟩, rcases p with ⟨x, y⟩, - simp only [is_bounded_bilinear_map_deriv_coe, prod.mk_sub_mk, h.map_sub_left, h.map_sub_right], - abel }, - have : (λ q : T, p - q.2) =o[𝓝 (p, p)] (λ q, (1:ℝ)), - from (is_o_one_iff _).2 (sub_self p ▸ tendsto_const_nhds.sub continuous_at_snd), - apply is_bounded_bilinear_map_apply.is_O_comp.trans_is_o, - refine is_o.trans_is_O _ (is_O_const_mul_self 1 _ _).of_norm_right, - refine is_o.mul_is_O _ (is_O_refl _ _), - exact (((h.is_bounded_linear_map_deriv.is_O_id ⊤).comp_tendsto le_top : _).trans_is_o - this).norm_left -end - -lemma is_bounded_bilinear_map.has_fderiv_at (h : is_bounded_bilinear_map 𝕜 b) (p : E × F) : - has_fderiv_at b (h.deriv p) p := -(h.has_strict_fderiv_at p).has_fderiv_at - -lemma is_bounded_bilinear_map.has_fderiv_within_at (h : is_bounded_bilinear_map 𝕜 b) (p : E × F) : - has_fderiv_within_at b (h.deriv p) u p := -(h.has_fderiv_at p).has_fderiv_within_at - -lemma is_bounded_bilinear_map.differentiable_at (h : is_bounded_bilinear_map 𝕜 b) (p : E × F) : - differentiable_at 𝕜 b p := -(h.has_fderiv_at p).differentiable_at - -lemma is_bounded_bilinear_map.differentiable_within_at (h : is_bounded_bilinear_map 𝕜 b) - (p : E × F) : - differentiable_within_at 𝕜 b u p := -(h.differentiable_at p).differentiable_within_at - -lemma is_bounded_bilinear_map.fderiv (h : is_bounded_bilinear_map 𝕜 b) (p : E × F) : - fderiv 𝕜 b p = h.deriv p := -has_fderiv_at.fderiv (h.has_fderiv_at p) - -lemma is_bounded_bilinear_map.fderiv_within (h : is_bounded_bilinear_map 𝕜 b) (p : E × F) - (hxs : unique_diff_within_at 𝕜 u p) : fderiv_within 𝕜 b u p = h.deriv p := -begin - rw differentiable_at.fderiv_within (h.differentiable_at p) hxs, - exact h.fderiv p -end - -lemma is_bounded_bilinear_map.differentiable (h : is_bounded_bilinear_map 𝕜 b) : - differentiable 𝕜 b := -λx, h.differentiable_at x - -lemma is_bounded_bilinear_map.differentiable_on (h : is_bounded_bilinear_map 𝕜 b) : - differentiable_on 𝕜 b u := -h.differentiable.differentiable_on - -variable (B : E →L[𝕜] F →L[𝕜] G) - -lemma continuous_linear_map.has_fderiv_within_at_of_bilinear - {f : G' → E} {g : G' → F} {f' : G' →L[𝕜] E} {g' : G' →L[𝕜] F} {x : G'} {s : set G'} - (hf : has_fderiv_within_at f f' s x) (hg : has_fderiv_within_at g g' s x) : - has_fderiv_within_at (λ y, B (f y) (g y)) (B.precompR G' (f x) g' + B.precompL G' f' (g x)) s x := -(B.is_bounded_bilinear_map.has_fderiv_at (f x, g x)).comp_has_fderiv_within_at x (hf.prod hg) - -lemma continuous_linear_map.has_fderiv_at_of_bilinear - {f : G' → E} {g : G' → F} {f' : G' →L[𝕜] E} {g' : G' →L[𝕜] F} {x : G'} - (hf : has_fderiv_at f f' x) (hg : has_fderiv_at g g' x) : - has_fderiv_at (λ y, B (f y) (g y)) (B.precompR G' (f x) g' + B.precompL G' f' (g x)) x := -(B.is_bounded_bilinear_map.has_fderiv_at (f x, g x)).comp x (hf.prod hg) - -lemma continuous_linear_map.fderiv_within_of_bilinear - {f : G' → E} {g : G' → F} {x : G'} {s : set G'} - (hf : differentiable_within_at 𝕜 f s x) (hg : differentiable_within_at 𝕜 g s x) - (hs : unique_diff_within_at 𝕜 s x) : - fderiv_within 𝕜 (λ y, B (f y) (g y)) s x = - (B.precompR G' (f x) (fderiv_within 𝕜 g s x) + B.precompL G' (fderiv_within 𝕜 f s x) (g x)) := -(B.has_fderiv_within_at_of_bilinear hf.has_fderiv_within_at hg.has_fderiv_within_at).fderiv_within - hs - -lemma continuous_linear_map.fderiv_of_bilinear {f : G' → E} {g : G' → F} {x : G'} - (hf : differentiable_at 𝕜 f x) (hg : differentiable_at 𝕜 g x) : - fderiv 𝕜 (λ y, B (f y) (g y)) x = - (B.precompR G' (f x) (fderiv 𝕜 g x) + B.precompL G' (fderiv 𝕜 f x) (g x)) := -(B.has_fderiv_at_of_bilinear hf.has_fderiv_at hg.has_fderiv_at).fderiv - -end bilinear_map - -section clm_comp_apply -/-! ### Derivative of the pointwise composition/application of continuous linear maps -/ - -variables {H : Type*} [normed_add_comm_group H] [normed_space 𝕜 H] {c : E → G →L[𝕜] H} - {c' : E →L[𝕜] G →L[𝕜] H} {d : E → F →L[𝕜] G} {d' : E →L[𝕜] F →L[𝕜] G} {u : E → G} - {u' : E →L[𝕜] G} - -lemma has_strict_fderiv_at.clm_comp (hc : has_strict_fderiv_at c c' x) - (hd : has_strict_fderiv_at d d' x) : has_strict_fderiv_at (λ y, (c y).comp (d y)) - ((compL 𝕜 F G H (c x)).comp d' + ((compL 𝕜 F G H).flip (d x)).comp c') x := -(is_bounded_bilinear_map_comp.has_strict_fderiv_at (c x, d x)).comp x $ hc.prod hd - -lemma has_fderiv_within_at.clm_comp (hc : has_fderiv_within_at c c' s x) - (hd : has_fderiv_within_at d d' s x) : has_fderiv_within_at (λ y, (c y).comp (d y)) - ((compL 𝕜 F G H (c x)).comp d' + ((compL 𝕜 F G H).flip (d x)).comp c') s x := -(is_bounded_bilinear_map_comp.has_fderiv_at (c x, d x)).comp_has_fderiv_within_at x $ hc.prod hd - -lemma has_fderiv_at.clm_comp (hc : has_fderiv_at c c' x) - (hd : has_fderiv_at d d' x) : has_fderiv_at (λ y, (c y).comp (d y)) - ((compL 𝕜 F G H (c x)).comp d' + ((compL 𝕜 F G H).flip (d x)).comp c') x := -(is_bounded_bilinear_map_comp.has_fderiv_at (c x, d x)).comp x $ hc.prod hd - -lemma differentiable_within_at.clm_comp - (hc : differentiable_within_at 𝕜 c s x) (hd : differentiable_within_at 𝕜 d s x) : - differentiable_within_at 𝕜 (λ y, (c y).comp (d y)) s x := -(hc.has_fderiv_within_at.clm_comp hd.has_fderiv_within_at).differentiable_within_at - -lemma differentiable_at.clm_comp (hc : differentiable_at 𝕜 c x) - (hd : differentiable_at 𝕜 d x) : differentiable_at 𝕜 (λ y, (c y).comp (d y)) x := -(hc.has_fderiv_at.clm_comp hd.has_fderiv_at).differentiable_at - -lemma differentiable_on.clm_comp (hc : differentiable_on 𝕜 c s) (hd : differentiable_on 𝕜 d s) : - differentiable_on 𝕜 (λ y, (c y).comp (d y)) s := -λx hx, (hc x hx).clm_comp (hd x hx) - -lemma differentiable.clm_comp (hc : differentiable 𝕜 c) (hd : differentiable 𝕜 d) : - differentiable 𝕜 (λ y, (c y).comp (d y)) := -λx, (hc x).clm_comp (hd x) - -lemma fderiv_within_clm_comp (hxs : unique_diff_within_at 𝕜 s x) - (hc : differentiable_within_at 𝕜 c s x) (hd : differentiable_within_at 𝕜 d s x) : - fderiv_within 𝕜 (λ y, (c y).comp (d y)) s x = - (compL 𝕜 F G H (c x)).comp (fderiv_within 𝕜 d s x) + - ((compL 𝕜 F G H).flip (d x)).comp (fderiv_within 𝕜 c s x) := -(hc.has_fderiv_within_at.clm_comp hd.has_fderiv_within_at).fderiv_within hxs - -lemma fderiv_clm_comp (hc : differentiable_at 𝕜 c x) (hd : differentiable_at 𝕜 d x) : - fderiv 𝕜 (λ y, (c y).comp (d y)) x = - (compL 𝕜 F G H (c x)).comp (fderiv 𝕜 d x) + - ((compL 𝕜 F G H).flip (d x)).comp (fderiv 𝕜 c x) := -(hc.has_fderiv_at.clm_comp hd.has_fderiv_at).fderiv - -lemma has_strict_fderiv_at.clm_apply (hc : has_strict_fderiv_at c c' x) - (hu : has_strict_fderiv_at u u' x) : - has_strict_fderiv_at (λ y, (c y) (u y)) ((c x).comp u' + c'.flip (u x)) x := -(is_bounded_bilinear_map_apply.has_strict_fderiv_at (c x, u x)).comp x (hc.prod hu) - -lemma has_fderiv_within_at.clm_apply (hc : has_fderiv_within_at c c' s x) - (hu : has_fderiv_within_at u u' s x) : - has_fderiv_within_at (λ y, (c y) (u y)) ((c x).comp u' + c'.flip (u x)) s x := -(is_bounded_bilinear_map_apply.has_fderiv_at (c x, u x)).comp_has_fderiv_within_at x (hc.prod hu) - -lemma has_fderiv_at.clm_apply (hc : has_fderiv_at c c' x) (hu : has_fderiv_at u u' x) : - has_fderiv_at (λ y, (c y) (u y)) ((c x).comp u' + c'.flip (u x)) x := -(is_bounded_bilinear_map_apply.has_fderiv_at (c x, u x)).comp x (hc.prod hu) - -lemma differentiable_within_at.clm_apply - (hc : differentiable_within_at 𝕜 c s x) (hu : differentiable_within_at 𝕜 u s x) : - differentiable_within_at 𝕜 (λ y, (c y) (u y)) s x := -(hc.has_fderiv_within_at.clm_apply hu.has_fderiv_within_at).differentiable_within_at - -lemma differentiable_at.clm_apply (hc : differentiable_at 𝕜 c x) - (hu : differentiable_at 𝕜 u x) : differentiable_at 𝕜 (λ y, (c y) (u y)) x := -(hc.has_fderiv_at.clm_apply hu.has_fderiv_at).differentiable_at - -lemma differentiable_on.clm_apply (hc : differentiable_on 𝕜 c s) (hu : differentiable_on 𝕜 u s) : - differentiable_on 𝕜 (λ y, (c y) (u y)) s := -λx hx, (hc x hx).clm_apply (hu x hx) - -lemma differentiable.clm_apply (hc : differentiable 𝕜 c) (hu : differentiable 𝕜 u) : - differentiable 𝕜 (λ y, (c y) (u y)) := -λx, (hc x).clm_apply (hu x) - -lemma fderiv_within_clm_apply (hxs : unique_diff_within_at 𝕜 s x) - (hc : differentiable_within_at 𝕜 c s x) (hu : differentiable_within_at 𝕜 u s x) : - fderiv_within 𝕜 (λ y, (c y) (u y)) s x = - ((c x).comp (fderiv_within 𝕜 u s x) + (fderiv_within 𝕜 c s x).flip (u x)) := -(hc.has_fderiv_within_at.clm_apply hu.has_fderiv_within_at).fderiv_within hxs - -lemma fderiv_clm_apply (hc : differentiable_at 𝕜 c x) (hu : differentiable_at 𝕜 u x) : - fderiv 𝕜 (λ y, (c y) (u y)) x = ((c x).comp (fderiv 𝕜 u x) + (fderiv 𝕜 c x).flip (u x)) := -(hc.has_fderiv_at.clm_apply hu.has_fderiv_at).fderiv - -end clm_comp_apply - -section smul -/-! ### Derivative of the product of a scalar-valued function and a vector-valued function - -If `c` is a differentiable scalar-valued function and `f` is a differentiable vector-valued -function, then `λ x, c x • f x` is differentiable as well. Lemmas in this section works for -function `c` taking values in the base field, as well as in a normed algebra over the base -field: e.g., they work for `c : E → ℂ` and `f : E → F` provided that `F` is a complex -normed vector space. --/ - -variables {𝕜' : Type*} [nontrivially_normed_field 𝕜'] [normed_algebra 𝕜 𝕜'] - [normed_space 𝕜' F] [is_scalar_tower 𝕜 𝕜' F] -variables {c : E → 𝕜'} {c' : E →L[𝕜] 𝕜'} - -theorem has_strict_fderiv_at.smul (hc : has_strict_fderiv_at c c' x) - (hf : has_strict_fderiv_at f f' x) : - has_strict_fderiv_at (λ y, c y • f y) (c x • f' + c'.smul_right (f x)) x := -(is_bounded_bilinear_map_smul.has_strict_fderiv_at (c x, f x)).comp x $ - hc.prod hf - -theorem has_fderiv_within_at.smul - (hc : has_fderiv_within_at c c' s x) (hf : has_fderiv_within_at f f' s x) : - has_fderiv_within_at (λ y, c y • f y) (c x • f' + c'.smul_right (f x)) s x := -(is_bounded_bilinear_map_smul.has_fderiv_at (c x, f x)).comp_has_fderiv_within_at x $ - hc.prod hf - -theorem has_fderiv_at.smul (hc : has_fderiv_at c c' x) (hf : has_fderiv_at f f' x) : - has_fderiv_at (λ y, c y • f y) (c x • f' + c'.smul_right (f x)) x := -(is_bounded_bilinear_map_smul.has_fderiv_at (c x, f x)).comp x $ - hc.prod hf - -lemma differentiable_within_at.smul - (hc : differentiable_within_at 𝕜 c s x) (hf : differentiable_within_at 𝕜 f s x) : - differentiable_within_at 𝕜 (λ y, c y • f y) s x := -(hc.has_fderiv_within_at.smul hf.has_fderiv_within_at).differentiable_within_at - -@[simp] lemma differentiable_at.smul (hc : differentiable_at 𝕜 c x) (hf : differentiable_at 𝕜 f x) : - differentiable_at 𝕜 (λ y, c y • f y) x := -(hc.has_fderiv_at.smul hf.has_fderiv_at).differentiable_at - -lemma differentiable_on.smul (hc : differentiable_on 𝕜 c s) (hf : differentiable_on 𝕜 f s) : - differentiable_on 𝕜 (λ y, c y • f y) s := -λx hx, (hc x hx).smul (hf x hx) - -@[simp] lemma differentiable.smul (hc : differentiable 𝕜 c) (hf : differentiable 𝕜 f) : - differentiable 𝕜 (λ y, c y • f y) := -λx, (hc x).smul (hf x) - -lemma fderiv_within_smul (hxs : unique_diff_within_at 𝕜 s x) - (hc : differentiable_within_at 𝕜 c s x) (hf : differentiable_within_at 𝕜 f s x) : - fderiv_within 𝕜 (λ y, c y • f y) s x = - c x • fderiv_within 𝕜 f s x + (fderiv_within 𝕜 c s x).smul_right (f x) := -(hc.has_fderiv_within_at.smul hf.has_fderiv_within_at).fderiv_within hxs - -lemma fderiv_smul (hc : differentiable_at 𝕜 c x) (hf : differentiable_at 𝕜 f x) : - fderiv 𝕜 (λ y, c y • f y) x = - c x • fderiv 𝕜 f x + (fderiv 𝕜 c x).smul_right (f x) := -(hc.has_fderiv_at.smul hf.has_fderiv_at).fderiv - -theorem has_strict_fderiv_at.smul_const (hc : has_strict_fderiv_at c c' x) (f : F) : - has_strict_fderiv_at (λ y, c y • f) (c'.smul_right f) x := -by simpa only [smul_zero, zero_add] using hc.smul (has_strict_fderiv_at_const f x) - -theorem has_fderiv_within_at.smul_const (hc : has_fderiv_within_at c c' s x) (f : F) : - has_fderiv_within_at (λ y, c y • f) (c'.smul_right f) s x := -by simpa only [smul_zero, zero_add] using hc.smul (has_fderiv_within_at_const f x s) - -theorem has_fderiv_at.smul_const (hc : has_fderiv_at c c' x) (f : F) : - has_fderiv_at (λ y, c y • f) (c'.smul_right f) x := -by simpa only [smul_zero, zero_add] using hc.smul (has_fderiv_at_const f x) - -lemma differentiable_within_at.smul_const - (hc : differentiable_within_at 𝕜 c s x) (f : F) : - differentiable_within_at 𝕜 (λ y, c y • f) s x := -(hc.has_fderiv_within_at.smul_const f).differentiable_within_at - -lemma differentiable_at.smul_const (hc : differentiable_at 𝕜 c x) (f : F) : - differentiable_at 𝕜 (λ y, c y • f) x := -(hc.has_fderiv_at.smul_const f).differentiable_at - -lemma differentiable_on.smul_const (hc : differentiable_on 𝕜 c s) (f : F) : - differentiable_on 𝕜 (λ y, c y • f) s := -λx hx, (hc x hx).smul_const f - -lemma differentiable.smul_const (hc : differentiable 𝕜 c) (f : F) : - differentiable 𝕜 (λ y, c y • f) := -λx, (hc x).smul_const f - -lemma fderiv_within_smul_const (hxs : unique_diff_within_at 𝕜 s x) - (hc : differentiable_within_at 𝕜 c s x) (f : F) : - fderiv_within 𝕜 (λ y, c y • f) s x = - (fderiv_within 𝕜 c s x).smul_right f := -(hc.has_fderiv_within_at.smul_const f).fderiv_within hxs - -lemma fderiv_smul_const (hc : differentiable_at 𝕜 c x) (f : F) : - fderiv 𝕜 (λ y, c y • f) x = (fderiv 𝕜 c x).smul_right f := -(hc.has_fderiv_at.smul_const f).fderiv - -end smul - -section mul -/-! ### Derivative of the product of two functions -/ - -variables {𝔸 𝔸' : Type*} [normed_ring 𝔸] [normed_comm_ring 𝔸'] [normed_algebra 𝕜 𝔸] - [normed_algebra 𝕜 𝔸'] {a b : E → 𝔸} {a' b' : E →L[𝕜] 𝔸} {c d : E → 𝔸'} {c' d' : E →L[𝕜] 𝔸'} - -theorem has_strict_fderiv_at.mul' {x : E} (ha : has_strict_fderiv_at a a' x) - (hb : has_strict_fderiv_at b b' x) : - has_strict_fderiv_at (λ y, a y * b y) (a x • b' + a'.smul_right (b x)) x := -((continuous_linear_map.mul 𝕜 𝔸).is_bounded_bilinear_map.has_strict_fderiv_at (a x, b x)).comp x - (ha.prod hb) - -theorem has_strict_fderiv_at.mul - (hc : has_strict_fderiv_at c c' x) (hd : has_strict_fderiv_at d d' x) : - has_strict_fderiv_at (λ y, c y * d y) (c x • d' + d x • c') x := -by { convert hc.mul' hd, ext z, apply mul_comm } - -theorem has_fderiv_within_at.mul' - (ha : has_fderiv_within_at a a' s x) (hb : has_fderiv_within_at b b' s x) : - has_fderiv_within_at (λ y, a y * b y) (a x • b' + a'.smul_right (b x)) s x := -((continuous_linear_map.mul 𝕜 𝔸).is_bounded_bilinear_map.has_fderiv_at - (a x, b x)).comp_has_fderiv_within_at x (ha.prod hb) - -theorem has_fderiv_within_at.mul - (hc : has_fderiv_within_at c c' s x) (hd : has_fderiv_within_at d d' s x) : - has_fderiv_within_at (λ y, c y * d y) (c x • d' + d x • c') s x := -by { convert hc.mul' hd, ext z, apply mul_comm } - -theorem has_fderiv_at.mul' - (ha : has_fderiv_at a a' x) (hb : has_fderiv_at b b' x) : - has_fderiv_at (λ y, a y * b y) (a x • b' + a'.smul_right (b x)) x := -((continuous_linear_map.mul 𝕜 𝔸).is_bounded_bilinear_map.has_fderiv_at (a x, b x)).comp x - (ha.prod hb) - -theorem has_fderiv_at.mul (hc : has_fderiv_at c c' x) (hd : has_fderiv_at d d' x) : - has_fderiv_at (λ y, c y * d y) (c x • d' + d x • c') x := -by { convert hc.mul' hd, ext z, apply mul_comm } - -lemma differentiable_within_at.mul - (ha : differentiable_within_at 𝕜 a s x) (hb : differentiable_within_at 𝕜 b s x) : - differentiable_within_at 𝕜 (λ y, a y * b y) s x := -(ha.has_fderiv_within_at.mul' hb.has_fderiv_within_at).differentiable_within_at - -@[simp] lemma differentiable_at.mul (ha : differentiable_at 𝕜 a x) (hb : differentiable_at 𝕜 b x) : - differentiable_at 𝕜 (λ y, a y * b y) x := -(ha.has_fderiv_at.mul' hb.has_fderiv_at).differentiable_at - -lemma differentiable_on.mul (ha : differentiable_on 𝕜 a s) (hb : differentiable_on 𝕜 b s) : - differentiable_on 𝕜 (λ y, a y * b y) s := -λx hx, (ha x hx).mul (hb x hx) - -@[simp] lemma differentiable.mul (ha : differentiable 𝕜 a) (hb : differentiable 𝕜 b) : - differentiable 𝕜 (λ y, a y * b y) := -λx, (ha x).mul (hb x) - -lemma differentiable_within_at.pow (ha : differentiable_within_at 𝕜 a s x) : - ∀ n : ℕ, differentiable_within_at 𝕜 (λ x, a x ^ n) s x -| 0 := by simp only [pow_zero, differentiable_within_at_const] -| (n + 1) := by simp only [pow_succ, differentiable_within_at.pow n, ha.mul] - -@[simp] lemma differentiable_at.pow (ha : differentiable_at 𝕜 a x) (n : ℕ) : - differentiable_at 𝕜 (λ x, a x ^ n) x := -differentiable_within_at_univ.mp $ ha.differentiable_within_at.pow n - -lemma differentiable_on.pow (ha : differentiable_on 𝕜 a s) (n : ℕ) : - differentiable_on 𝕜 (λ x, a x ^ n) s := -λ x h, (ha x h).pow n - -@[simp] lemma differentiable.pow (ha : differentiable 𝕜 a) (n : ℕ) : - differentiable 𝕜 (λ x, a x ^ n) := -λx, (ha x).pow n - -lemma fderiv_within_mul' (hxs : unique_diff_within_at 𝕜 s x) - (ha : differentiable_within_at 𝕜 a s x) (hb : differentiable_within_at 𝕜 b s x) : - fderiv_within 𝕜 (λ y, a y * b y) s x = - a x • fderiv_within 𝕜 b s x + (fderiv_within 𝕜 a s x).smul_right (b x) := -(ha.has_fderiv_within_at.mul' hb.has_fderiv_within_at).fderiv_within hxs - -lemma fderiv_within_mul (hxs : unique_diff_within_at 𝕜 s x) - (hc : differentiable_within_at 𝕜 c s x) (hd : differentiable_within_at 𝕜 d s x) : - fderiv_within 𝕜 (λ y, c y * d y) s x = - c x • fderiv_within 𝕜 d s x + d x • fderiv_within 𝕜 c s x := -(hc.has_fderiv_within_at.mul hd.has_fderiv_within_at).fderiv_within hxs - -lemma fderiv_mul' (ha : differentiable_at 𝕜 a x) (hb : differentiable_at 𝕜 b x) : - fderiv 𝕜 (λ y, a y * b y) x = - a x • fderiv 𝕜 b x + (fderiv 𝕜 a x).smul_right (b x) := -(ha.has_fderiv_at.mul' hb.has_fderiv_at).fderiv - -lemma fderiv_mul (hc : differentiable_at 𝕜 c x) (hd : differentiable_at 𝕜 d x) : - fderiv 𝕜 (λ y, c y * d y) x = - c x • fderiv 𝕜 d x + d x • fderiv 𝕜 c x := -(hc.has_fderiv_at.mul hd.has_fderiv_at).fderiv - -theorem has_strict_fderiv_at.mul_const' (ha : has_strict_fderiv_at a a' x) (b : 𝔸) : - has_strict_fderiv_at (λ y, a y * b) (a'.smul_right b) x := -(((continuous_linear_map.mul 𝕜 𝔸).flip b).has_strict_fderiv_at).comp x ha - -theorem has_strict_fderiv_at.mul_const (hc : has_strict_fderiv_at c c' x) (d : 𝔸') : - has_strict_fderiv_at (λ y, c y * d) (d • c') x := -by { convert hc.mul_const' d, ext z, apply mul_comm } - -theorem has_fderiv_within_at.mul_const' (ha : has_fderiv_within_at a a' s x) (b : 𝔸) : - has_fderiv_within_at (λ y, a y * b) (a'.smul_right b) s x := -(((continuous_linear_map.mul 𝕜 𝔸).flip b).has_fderiv_at).comp_has_fderiv_within_at x ha - -theorem has_fderiv_within_at.mul_const (hc : has_fderiv_within_at c c' s x) (d : 𝔸') : - has_fderiv_within_at (λ y, c y * d) (d • c') s x := -by { convert hc.mul_const' d, ext z, apply mul_comm } - -theorem has_fderiv_at.mul_const' (ha : has_fderiv_at a a' x) (b : 𝔸) : - has_fderiv_at (λ y, a y * b) (a'.smul_right b) x := -(((continuous_linear_map.mul 𝕜 𝔸).flip b).has_fderiv_at).comp x ha - -theorem has_fderiv_at.mul_const (hc : has_fderiv_at c c' x) (d : 𝔸') : - has_fderiv_at (λ y, c y * d) (d • c') x := -by { convert hc.mul_const' d, ext z, apply mul_comm } - -lemma differentiable_within_at.mul_const - (ha : differentiable_within_at 𝕜 a s x) (b : 𝔸) : - differentiable_within_at 𝕜 (λ y, a y * b) s x := -(ha.has_fderiv_within_at.mul_const' b).differentiable_within_at - -lemma differentiable_at.mul_const (ha : differentiable_at 𝕜 a x) (b : 𝔸) : - differentiable_at 𝕜 (λ y, a y * b) x := -(ha.has_fderiv_at.mul_const' b).differentiable_at - -lemma differentiable_on.mul_const (ha : differentiable_on 𝕜 a s) (b : 𝔸) : - differentiable_on 𝕜 (λ y, a y * b) s := -λx hx, (ha x hx).mul_const b - -lemma differentiable.mul_const (ha : differentiable 𝕜 a) (b : 𝔸) : - differentiable 𝕜 (λ y, a y * b) := -λx, (ha x).mul_const b - -lemma fderiv_within_mul_const' (hxs : unique_diff_within_at 𝕜 s x) - (ha : differentiable_within_at 𝕜 a s x) (b : 𝔸) : - fderiv_within 𝕜 (λ y, a y * b) s x = (fderiv_within 𝕜 a s x).smul_right b := -(ha.has_fderiv_within_at.mul_const' b).fderiv_within hxs - -lemma fderiv_within_mul_const (hxs : unique_diff_within_at 𝕜 s x) - (hc : differentiable_within_at 𝕜 c s x) (d : 𝔸') : - fderiv_within 𝕜 (λ y, c y * d) s x = d • fderiv_within 𝕜 c s x := -(hc.has_fderiv_within_at.mul_const d).fderiv_within hxs - -lemma fderiv_mul_const' (ha : differentiable_at 𝕜 a x) (b : 𝔸) : - fderiv 𝕜 (λ y, a y * b) x = (fderiv 𝕜 a x).smul_right b := -(ha.has_fderiv_at.mul_const' b).fderiv - -lemma fderiv_mul_const (hc : differentiable_at 𝕜 c x) (d : 𝔸') : - fderiv 𝕜 (λ y, c y * d) x = d • fderiv 𝕜 c x := -(hc.has_fderiv_at.mul_const d).fderiv - -theorem has_strict_fderiv_at.const_mul (ha : has_strict_fderiv_at a a' x) (b : 𝔸) : - has_strict_fderiv_at (λ y, b * a y) (b • a') x := -(((continuous_linear_map.mul 𝕜 𝔸) b).has_strict_fderiv_at).comp x ha - -theorem has_fderiv_within_at.const_mul - (ha : has_fderiv_within_at a a' s x) (b : 𝔸) : - has_fderiv_within_at (λ y, b * a y) (b • a') s x := -(((continuous_linear_map.mul 𝕜 𝔸) b).has_fderiv_at).comp_has_fderiv_within_at x ha - -theorem has_fderiv_at.const_mul (ha : has_fderiv_at a a' x) (b : 𝔸) : - has_fderiv_at (λ y, b * a y) (b • a') x := -(((continuous_linear_map.mul 𝕜 𝔸) b).has_fderiv_at).comp x ha - -lemma differentiable_within_at.const_mul - (ha : differentiable_within_at 𝕜 a s x) (b : 𝔸) : - differentiable_within_at 𝕜 (λ y, b * a y) s x := -(ha.has_fderiv_within_at.const_mul b).differentiable_within_at - -lemma differentiable_at.const_mul (ha : differentiable_at 𝕜 a x) (b : 𝔸) : - differentiable_at 𝕜 (λ y, b * a y) x := -(ha.has_fderiv_at.const_mul b).differentiable_at - -lemma differentiable_on.const_mul (ha : differentiable_on 𝕜 a s) (b : 𝔸) : - differentiable_on 𝕜 (λ y, b * a y) s := -λx hx, (ha x hx).const_mul b - -lemma differentiable.const_mul (ha : differentiable 𝕜 a) (b : 𝔸) : - differentiable 𝕜 (λ y, b * a y) := -λx, (ha x).const_mul b - -lemma fderiv_within_const_mul (hxs : unique_diff_within_at 𝕜 s x) - (ha : differentiable_within_at 𝕜 a s x) (b : 𝔸) : - fderiv_within 𝕜 (λ y, b * a y) s x = b • fderiv_within 𝕜 a s x := -(ha.has_fderiv_within_at.const_mul b).fderiv_within hxs - -lemma fderiv_const_mul (ha : differentiable_at 𝕜 a x) (b : 𝔸) : - fderiv 𝕜 (λ y, b * a y) x = b • fderiv 𝕜 a x := -(ha.has_fderiv_at.const_mul b).fderiv - -end mul - -section algebra_inverse -variables {R : Type*} [normed_ring R] [normed_algebra 𝕜 R] [complete_space R] -open normed_ring continuous_linear_map ring - -/-- At an invertible element `x` of a normed algebra `R`, the Fréchet derivative of the inversion -operation is the linear map `λ t, - x⁻¹ * t * x⁻¹`. -/ -lemma has_fderiv_at_ring_inverse (x : Rˣ) : - has_fderiv_at ring.inverse (-mul_left_right 𝕜 R ↑x⁻¹ ↑x⁻¹) x := -begin - have h_is_o : (λ (t : R), inverse (↑x + t) - ↑x⁻¹ + ↑x⁻¹ * t * ↑x⁻¹) =o[𝓝 0] (λ (t : R), t), - { refine (inverse_add_norm_diff_second_order x).trans_is_o ((is_o_norm_norm).mp _), - simp only [norm_pow, norm_norm], - have h12 : 1 < 2 := by norm_num, - convert (asymptotics.is_o_pow_pow h12).comp_tendsto tendsto_norm_zero, - ext, simp }, - have h_lim : tendsto (λ (y:R), y - x) (𝓝 x) (𝓝 0), - { refine tendsto_zero_iff_norm_tendsto_zero.mpr _, - exact tendsto_iff_norm_tendsto_zero.mp tendsto_id }, - simp only [has_fderiv_at, has_fderiv_at_filter], - convert h_is_o.comp_tendsto h_lim, - ext y, - simp only [coe_comp', function.comp_app, mul_left_right_apply, neg_apply, inverse_unit x, - units.inv_mul, add_sub_cancel'_right, mul_sub, sub_mul, one_mul, sub_neg_eq_add] -end - -lemma differentiable_at_inverse (x : Rˣ) : differentiable_at 𝕜 (@ring.inverse R _) x := -(has_fderiv_at_ring_inverse x).differentiable_at - -lemma fderiv_inverse (x : Rˣ) : - fderiv 𝕜 (@ring.inverse R _) x = - mul_left_right 𝕜 R ↑x⁻¹ ↑x⁻¹ := -(has_fderiv_at_ring_inverse x).fderiv - -end algebra_inverse - -namespace continuous_linear_equiv -/-! ### Differentiability of linear equivs, and invariance of differentiability -/ - -variable (iso : E ≃L[𝕜] F) - -protected lemma has_strict_fderiv_at : - has_strict_fderiv_at iso (iso : E →L[𝕜] F) x := -iso.to_continuous_linear_map.has_strict_fderiv_at - -protected lemma has_fderiv_within_at : - has_fderiv_within_at iso (iso : E →L[𝕜] F) s x := -iso.to_continuous_linear_map.has_fderiv_within_at - -protected lemma has_fderiv_at : has_fderiv_at iso (iso : E →L[𝕜] F) x := -iso.to_continuous_linear_map.has_fderiv_at_filter - -protected lemma differentiable_at : differentiable_at 𝕜 iso x := -iso.has_fderiv_at.differentiable_at - -protected lemma differentiable_within_at : - differentiable_within_at 𝕜 iso s x := -iso.differentiable_at.differentiable_within_at - -protected lemma fderiv : fderiv 𝕜 iso x = iso := -iso.has_fderiv_at.fderiv - -protected lemma fderiv_within (hxs : unique_diff_within_at 𝕜 s x) : - fderiv_within 𝕜 iso s x = iso := -iso.to_continuous_linear_map.fderiv_within hxs - -protected lemma differentiable : differentiable 𝕜 iso := -λx, iso.differentiable_at - -protected lemma differentiable_on : differentiable_on 𝕜 iso s := -iso.differentiable.differentiable_on - -lemma comp_differentiable_within_at_iff {f : G → E} {s : set G} {x : G} : - differentiable_within_at 𝕜 (iso ∘ f) s x ↔ differentiable_within_at 𝕜 f s x := -begin - refine ⟨λ H, _, λ H, iso.differentiable.differentiable_at.comp_differentiable_within_at x H⟩, - have : differentiable_within_at 𝕜 (iso.symm ∘ (iso ∘ f)) s x := - iso.symm.differentiable.differentiable_at.comp_differentiable_within_at x H, - rwa [← function.comp.assoc iso.symm iso f, iso.symm_comp_self] at this, -end - -lemma comp_differentiable_at_iff {f : G → E} {x : G} : - differentiable_at 𝕜 (iso ∘ f) x ↔ differentiable_at 𝕜 f x := -by rw [← differentiable_within_at_univ, ← differentiable_within_at_univ, - iso.comp_differentiable_within_at_iff] - -lemma comp_differentiable_on_iff {f : G → E} {s : set G} : - differentiable_on 𝕜 (iso ∘ f) s ↔ differentiable_on 𝕜 f s := -begin - rw [differentiable_on, differentiable_on], - simp only [iso.comp_differentiable_within_at_iff], -end - -lemma comp_differentiable_iff {f : G → E} : - differentiable 𝕜 (iso ∘ f) ↔ differentiable 𝕜 f := -begin - rw [← differentiable_on_univ, ← differentiable_on_univ], - exact iso.comp_differentiable_on_iff -end - -lemma comp_has_fderiv_within_at_iff - {f : G → E} {s : set G} {x : G} {f' : G →L[𝕜] E} : - has_fderiv_within_at (iso ∘ f) ((iso : E →L[𝕜] F).comp f') s x ↔ has_fderiv_within_at f f' s x := -begin - refine ⟨λ H, _, λ H, iso.has_fderiv_at.comp_has_fderiv_within_at x H⟩, - have A : f = iso.symm ∘ (iso ∘ f), by { rw [← function.comp.assoc, iso.symm_comp_self], refl }, - have B : f' = (iso.symm : F →L[𝕜] E).comp ((iso : E →L[𝕜] F).comp f'), - by rw [← continuous_linear_map.comp_assoc, iso.coe_symm_comp_coe, - continuous_linear_map.id_comp], - rw [A, B], - exact iso.symm.has_fderiv_at.comp_has_fderiv_within_at x H -end - -lemma comp_has_strict_fderiv_at_iff {f : G → E} {x : G} {f' : G →L[𝕜] E} : - has_strict_fderiv_at (iso ∘ f) ((iso : E →L[𝕜] F).comp f') x ↔ has_strict_fderiv_at f f' x := -begin - refine ⟨λ H, _, λ H, iso.has_strict_fderiv_at.comp x H⟩, - convert iso.symm.has_strict_fderiv_at.comp x H; ext z; apply (iso.symm_apply_apply _).symm -end - -lemma comp_has_fderiv_at_iff {f : G → E} {x : G} {f' : G →L[𝕜] E} : - has_fderiv_at (iso ∘ f) ((iso : E →L[𝕜] F).comp f') x ↔ has_fderiv_at f f' x := -by simp_rw [← has_fderiv_within_at_univ, iso.comp_has_fderiv_within_at_iff] - -lemma comp_has_fderiv_within_at_iff' - {f : G → E} {s : set G} {x : G} {f' : G →L[𝕜] F} : - has_fderiv_within_at (iso ∘ f) f' s x ↔ - has_fderiv_within_at f ((iso.symm : F →L[𝕜] E).comp f') s x := -by rw [← iso.comp_has_fderiv_within_at_iff, ← continuous_linear_map.comp_assoc, - iso.coe_comp_coe_symm, continuous_linear_map.id_comp] - -lemma comp_has_fderiv_at_iff' {f : G → E} {x : G} {f' : G →L[𝕜] F} : - has_fderiv_at (iso ∘ f) f' x ↔ has_fderiv_at f ((iso.symm : F →L[𝕜] E).comp f') x := -by simp_rw [← has_fderiv_within_at_univ, iso.comp_has_fderiv_within_at_iff'] - -lemma comp_fderiv_within {f : G → E} {s : set G} {x : G} - (hxs : unique_diff_within_at 𝕜 s x) : - fderiv_within 𝕜 (iso ∘ f) s x = (iso : E →L[𝕜] F).comp (fderiv_within 𝕜 f s x) := -begin - by_cases h : differentiable_within_at 𝕜 f s x, - { rw [fderiv.comp_fderiv_within x iso.differentiable_at h hxs, iso.fderiv] }, - { have : ¬differentiable_within_at 𝕜 (iso ∘ f) s x, - from mt iso.comp_differentiable_within_at_iff.1 h, - rw [fderiv_within_zero_of_not_differentiable_within_at h, - fderiv_within_zero_of_not_differentiable_within_at this, - continuous_linear_map.comp_zero] } -end - -lemma comp_fderiv {f : G → E} {x : G} : - fderiv 𝕜 (iso ∘ f) x = (iso : E →L[𝕜] F).comp (fderiv 𝕜 f x) := -begin - rw [← fderiv_within_univ, ← fderiv_within_univ], - exact iso.comp_fderiv_within unique_diff_within_at_univ, -end - -lemma comp_right_differentiable_within_at_iff {f : F → G} {s : set F} {x : E} : - differentiable_within_at 𝕜 (f ∘ iso) (iso ⁻¹' s) x ↔ differentiable_within_at 𝕜 f s (iso x) := -begin - refine ⟨λ H, _, λ H, H.comp x iso.differentiable_within_at (maps_to_preimage _ s)⟩, - have : differentiable_within_at 𝕜 ((f ∘ iso) ∘ iso.symm) s (iso x), - { rw ← iso.symm_apply_apply x at H, - apply H.comp (iso x) iso.symm.differentiable_within_at, - assume y hy, - simpa only [mem_preimage, apply_symm_apply] using hy }, - rwa [function.comp.assoc, iso.self_comp_symm] at this, -end - -lemma comp_right_differentiable_at_iff {f : F → G} {x : E} : - differentiable_at 𝕜 (f ∘ iso) x ↔ differentiable_at 𝕜 f (iso x) := -by simp only [← differentiable_within_at_univ, ← iso.comp_right_differentiable_within_at_iff, - preimage_univ] - -lemma comp_right_differentiable_on_iff {f : F → G} {s : set F} : - differentiable_on 𝕜 (f ∘ iso) (iso ⁻¹' s) ↔ differentiable_on 𝕜 f s := -begin - refine ⟨λ H y hy, _, λ H y hy, iso.comp_right_differentiable_within_at_iff.2 (H _ hy)⟩, - rw [← iso.apply_symm_apply y, ← comp_right_differentiable_within_at_iff], - apply H, - simpa only [mem_preimage, apply_symm_apply] using hy, -end - -lemma comp_right_differentiable_iff {f : F → G} : - differentiable 𝕜 (f ∘ iso) ↔ differentiable 𝕜 f := -by simp only [← differentiable_on_univ, ← iso.comp_right_differentiable_on_iff, preimage_univ] - -lemma comp_right_has_fderiv_within_at_iff - {f : F → G} {s : set F} {x : E} {f' : F →L[𝕜] G} : - has_fderiv_within_at (f ∘ iso) (f'.comp (iso : E →L[𝕜] F)) (iso ⁻¹' s) x ↔ - has_fderiv_within_at f f' s (iso x) := -begin - refine ⟨λ H, _, λ H, H.comp x iso.has_fderiv_within_at (maps_to_preimage _ s)⟩, - rw [← iso.symm_apply_apply x] at H, - have A : f = (f ∘ iso) ∘ iso.symm, by { rw [function.comp.assoc, iso.self_comp_symm], refl }, - have B : f' = (f'.comp (iso : E →L[𝕜] F)).comp (iso.symm : F →L[𝕜] E), - by rw [continuous_linear_map.comp_assoc, iso.coe_comp_coe_symm, - continuous_linear_map.comp_id], - rw [A, B], - apply H.comp (iso x) iso.symm.has_fderiv_within_at, - assume y hy, - simpa only [mem_preimage, apply_symm_apply] using hy -end - -lemma comp_right_has_fderiv_at_iff {f : F → G} {x : E} {f' : F →L[𝕜] G} : - has_fderiv_at (f ∘ iso) (f'.comp (iso : E →L[𝕜] F)) x ↔ has_fderiv_at f f' (iso x) := -by simp only [← has_fderiv_within_at_univ, ← comp_right_has_fderiv_within_at_iff, preimage_univ] - -lemma comp_right_has_fderiv_within_at_iff' - {f : F → G} {s : set F} {x : E} {f' : E →L[𝕜] G} : - has_fderiv_within_at (f ∘ iso) f' (iso ⁻¹' s) x ↔ - has_fderiv_within_at f (f'.comp (iso.symm : F →L[𝕜] E)) s (iso x) := -by rw [← iso.comp_right_has_fderiv_within_at_iff, continuous_linear_map.comp_assoc, - iso.coe_symm_comp_coe, continuous_linear_map.comp_id] - -lemma comp_right_has_fderiv_at_iff' {f : F → G} {x : E} {f' : E →L[𝕜] G} : - has_fderiv_at (f ∘ iso) f' x ↔ has_fderiv_at f (f'.comp (iso.symm : F →L[𝕜] E)) (iso x) := -by simp only [← has_fderiv_within_at_univ, ← iso.comp_right_has_fderiv_within_at_iff', - preimage_univ] - -lemma comp_right_fderiv_within {f : F → G} {s : set F} {x : E} - (hxs : unique_diff_within_at 𝕜 (iso ⁻¹' s) x) : - fderiv_within 𝕜 (f ∘ iso) (iso ⁻¹'s) x = (fderiv_within 𝕜 f s (iso x)).comp (iso : E →L[𝕜] F) := -begin - by_cases h : differentiable_within_at 𝕜 f s (iso x), - { exact (iso.comp_right_has_fderiv_within_at_iff.2 (h.has_fderiv_within_at)).fderiv_within hxs }, - { have : ¬ differentiable_within_at 𝕜 (f ∘ iso) (iso ⁻¹' s) x, - { assume h', exact h (iso.comp_right_differentiable_within_at_iff.1 h') }, - rw [fderiv_within_zero_of_not_differentiable_within_at h, - fderiv_within_zero_of_not_differentiable_within_at this, continuous_linear_map.zero_comp] } -end - -lemma comp_right_fderiv {f : F → G} {x : E} : - fderiv 𝕜 (f ∘ iso) x = (fderiv 𝕜 f (iso x)).comp (iso : E →L[𝕜] F) := -begin - rw [← fderiv_within_univ, ← fderiv_within_univ, ← iso.comp_right_fderiv_within, preimage_univ], - exact unique_diff_within_at_univ, -end - -end continuous_linear_equiv - -namespace linear_isometry_equiv -/-! ### Differentiability of linear isometry equivs, and invariance of differentiability -/ - -variable (iso : E ≃ₗᵢ[𝕜] F) - -protected lemma has_strict_fderiv_at : has_strict_fderiv_at iso (iso : E →L[𝕜] F) x := -(iso : E ≃L[𝕜] F).has_strict_fderiv_at - -protected lemma has_fderiv_within_at : has_fderiv_within_at iso (iso : E →L[𝕜] F) s x := -(iso : E ≃L[𝕜] F).has_fderiv_within_at - -protected lemma has_fderiv_at : has_fderiv_at iso (iso : E →L[𝕜] F) x := -(iso : E ≃L[𝕜] F).has_fderiv_at - -protected lemma differentiable_at : differentiable_at 𝕜 iso x := -iso.has_fderiv_at.differentiable_at - -protected lemma differentiable_within_at : - differentiable_within_at 𝕜 iso s x := -iso.differentiable_at.differentiable_within_at - -protected lemma fderiv : fderiv 𝕜 iso x = iso := iso.has_fderiv_at.fderiv - -protected lemma fderiv_within (hxs : unique_diff_within_at 𝕜 s x) : - fderiv_within 𝕜 iso s x = iso := -(iso : E ≃L[𝕜] F).fderiv_within hxs - -protected lemma differentiable : differentiable 𝕜 iso := -λx, iso.differentiable_at - -protected lemma differentiable_on : differentiable_on 𝕜 iso s := -iso.differentiable.differentiable_on - -lemma comp_differentiable_within_at_iff {f : G → E} {s : set G} {x : G} : - differentiable_within_at 𝕜 (iso ∘ f) s x ↔ differentiable_within_at 𝕜 f s x := -(iso : E ≃L[𝕜] F).comp_differentiable_within_at_iff - -lemma comp_differentiable_at_iff {f : G → E} {x : G} : - differentiable_at 𝕜 (iso ∘ f) x ↔ differentiable_at 𝕜 f x := -(iso : E ≃L[𝕜] F).comp_differentiable_at_iff - -lemma comp_differentiable_on_iff {f : G → E} {s : set G} : - differentiable_on 𝕜 (iso ∘ f) s ↔ differentiable_on 𝕜 f s := -(iso : E ≃L[𝕜] F).comp_differentiable_on_iff - -lemma comp_differentiable_iff {f : G → E} : - differentiable 𝕜 (iso ∘ f) ↔ differentiable 𝕜 f := -(iso : E ≃L[𝕜] F).comp_differentiable_iff - -lemma comp_has_fderiv_within_at_iff - {f : G → E} {s : set G} {x : G} {f' : G →L[𝕜] E} : - has_fderiv_within_at (iso ∘ f) ((iso : E →L[𝕜] F).comp f') s x ↔ has_fderiv_within_at f f' s x := -(iso : E ≃L[𝕜] F).comp_has_fderiv_within_at_iff - -lemma comp_has_strict_fderiv_at_iff {f : G → E} {x : G} {f' : G →L[𝕜] E} : - has_strict_fderiv_at (iso ∘ f) ((iso : E →L[𝕜] F).comp f') x ↔ has_strict_fderiv_at f f' x := -(iso : E ≃L[𝕜] F).comp_has_strict_fderiv_at_iff - -lemma comp_has_fderiv_at_iff {f : G → E} {x : G} {f' : G →L[𝕜] E} : - has_fderiv_at (iso ∘ f) ((iso : E →L[𝕜] F).comp f') x ↔ has_fderiv_at f f' x := -(iso : E ≃L[𝕜] F).comp_has_fderiv_at_iff - -lemma comp_has_fderiv_within_at_iff' - {f : G → E} {s : set G} {x : G} {f' : G →L[𝕜] F} : - has_fderiv_within_at (iso ∘ f) f' s x ↔ - has_fderiv_within_at f ((iso.symm : F →L[𝕜] E).comp f') s x := -(iso : E ≃L[𝕜] F).comp_has_fderiv_within_at_iff' - -lemma comp_has_fderiv_at_iff' {f : G → E} {x : G} {f' : G →L[𝕜] F} : - has_fderiv_at (iso ∘ f) f' x ↔ has_fderiv_at f ((iso.symm : F →L[𝕜] E).comp f') x := -(iso : E ≃L[𝕜] F).comp_has_fderiv_at_iff' - -lemma comp_fderiv_within {f : G → E} {s : set G} {x : G} - (hxs : unique_diff_within_at 𝕜 s x) : - fderiv_within 𝕜 (iso ∘ f) s x = (iso : E →L[𝕜] F).comp (fderiv_within 𝕜 f s x) := -(iso : E ≃L[𝕜] F).comp_fderiv_within hxs - -lemma comp_fderiv {f : G → E} {x : G} : - fderiv 𝕜 (iso ∘ f) x = (iso : E →L[𝕜] F).comp (fderiv 𝕜 f x) := -(iso : E ≃L[𝕜] F).comp_fderiv - -end linear_isometry_equiv - -/-- If `f (g y) = y` for `y` in some neighborhood of `a`, `g` is continuous at `a`, and `f` has an -invertible derivative `f'` at `g a` in the strict sense, then `g` has the derivative `f'⁻¹` at `a` -in the strict sense. - -This is one of the easy parts of the inverse function theorem: it assumes that we already have an -inverse function. -/ -theorem has_strict_fderiv_at.of_local_left_inverse {f : E → F} {f' : E ≃L[𝕜] F} {g : F → E} {a : F} - (hg : continuous_at g a) (hf : has_strict_fderiv_at f (f' : E →L[𝕜] F) (g a)) - (hfg : ∀ᶠ y in 𝓝 a, f (g y) = y) : - has_strict_fderiv_at g (f'.symm : F →L[𝕜] E) a := -begin - replace hg := hg.prod_map' hg, - replace hfg := hfg.prod_mk_nhds hfg, - have : (λ p : F × F, g p.1 - g p.2 - f'.symm (p.1 - p.2)) =O[𝓝 (a, a)] - (λ p : F × F, f' (g p.1 - g p.2) - (p.1 - p.2)), - { refine ((f'.symm : F →L[𝕜] E).is_O_comp _ _).congr (λ x, _) (λ _, rfl), - simp }, - refine this.trans_is_o _, clear this, - refine ((hf.comp_tendsto hg).symm.congr' (hfg.mono _) - (eventually_of_forall $ λ _, rfl)).trans_is_O _, - { rintros p ⟨hp1, hp2⟩, - simp [hp1, hp2] }, - { refine (hf.is_O_sub_rev.comp_tendsto hg).congr' - (eventually_of_forall $ λ _, rfl) (hfg.mono _), - rintros p ⟨hp1, hp2⟩, - simp only [(∘), hp1, hp2] } -end - -/-- If `f (g y) = y` for `y` in some neighborhood of `a`, `g` is continuous at `a`, and `f` has an -invertible derivative `f'` at `g a`, then `g` has the derivative `f'⁻¹` at `a`. - -This is one of the easy parts of the inverse function theorem: it assumes that we already have -an inverse function. -/ -theorem has_fderiv_at.of_local_left_inverse {f : E → F} {f' : E ≃L[𝕜] F} {g : F → E} {a : F} - (hg : continuous_at g a) (hf : has_fderiv_at f (f' : E →L[𝕜] F) (g a)) - (hfg : ∀ᶠ y in 𝓝 a, f (g y) = y) : - has_fderiv_at g (f'.symm : F →L[𝕜] E) a := -begin - have : (λ x : F, g x - g a - f'.symm (x - a)) =O[𝓝 a] (λ x : F, f' (g x - g a) - (x - a)), - { refine ((f'.symm : F →L[𝕜] E).is_O_comp _ _).congr (λ x, _) (λ _, rfl), - simp }, - refine this.trans_is_o _, clear this, - refine ((hf.comp_tendsto hg).symm.congr' (hfg.mono _) - (eventually_of_forall $ λ _, rfl)).trans_is_O _, - { rintros p hp, - simp [hp, hfg.self_of_nhds] }, - { refine ((hf.is_O_sub_rev f'.antilipschitz).comp_tendsto hg).congr' - (eventually_of_forall $ λ _, rfl) (hfg.mono _), - rintros p hp, - simp only [(∘), hp, hfg.self_of_nhds] } -end - -/-- If `f` is a local homeomorphism defined on a neighbourhood of `f.symm a`, and `f` has an -invertible derivative `f'` in the sense of strict differentiability at `f.symm a`, then `f.symm` has -the derivative `f'⁻¹` at `a`. - -This is one of the easy parts of the inverse function theorem: it assumes that we already have -an inverse function. -/ -lemma local_homeomorph.has_strict_fderiv_at_symm (f : local_homeomorph E F) {f' : E ≃L[𝕜] F} {a : F} - (ha : a ∈ f.target) (htff' : has_strict_fderiv_at f (f' : E →L[𝕜] F) (f.symm a)) : - has_strict_fderiv_at f.symm (f'.symm : F →L[𝕜] E) a := -htff'.of_local_left_inverse (f.symm.continuous_at ha) (f.eventually_right_inverse ha) - -/-- If `f` is a local homeomorphism defined on a neighbourhood of `f.symm a`, and `f` has an -invertible derivative `f'` at `f.symm a`, then `f.symm` has the derivative `f'⁻¹` at `a`. - -This is one of the easy parts of the inverse function theorem: it assumes that we already have -an inverse function. -/ -lemma local_homeomorph.has_fderiv_at_symm (f : local_homeomorph E F) {f' : E ≃L[𝕜] F} {a : F} - (ha : a ∈ f.target) (htff' : has_fderiv_at f (f' : E →L[𝕜] F) (f.symm a)) : - has_fderiv_at f.symm (f'.symm : F →L[𝕜] E) a := -htff'.of_local_left_inverse (f.symm.continuous_at ha) (f.eventually_right_inverse ha) - -lemma has_fderiv_within_at.eventually_ne (h : has_fderiv_within_at f f' s x) - (hf' : ∃ C, ∀ z, ‖z‖ ≤ C * ‖f' z‖) : - ∀ᶠ z in 𝓝[s \ {x}] x, f z ≠ f x := -begin - rw [nhds_within, diff_eq, ← inf_principal, ← inf_assoc, eventually_inf_principal], - have A : (λ z, z - x) =O[𝓝[s] x] (λ z, f' (z - x)) := - (is_O_iff.2 $ hf'.imp $ λ C hC, eventually_of_forall $ λ z, hC _), - have : (λ z, f z - f x) ~[𝓝[s] x] (λ z, f' (z - x)) := h.trans_is_O A, - simpa [not_imp_not, sub_eq_zero] using (A.trans this.is_O_symm).eq_zero_imp -end - -lemma has_fderiv_at.eventually_ne (h : has_fderiv_at f f' x) (hf' : ∃ C, ∀ z, ‖z‖ ≤ C * ‖f' z‖) : - ∀ᶠ z in 𝓝[≠] x, f z ≠ f x := -by simpa only [compl_eq_univ_diff] using (has_fderiv_within_at_univ.2 h).eventually_ne hf' - -end - -section -/- - In the special case of a normed space over the reals, - we can use scalar multiplication in the `tendsto` characterization - of the Fréchet derivative. --/ - - -variables {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] -variables {F : Type*} [normed_add_comm_group F] [normed_space ℝ F] -variables {f : E → F} {f' : E →L[ℝ] F} {x : E} - -theorem has_fderiv_at_filter_real_equiv {L : filter E} : - tendsto (λ x' : E, ‖x' - x‖⁻¹ * ‖f x' - f x - f' (x' - x)‖) L (𝓝 0) ↔ - tendsto (λ x' : E, ‖x' - x‖⁻¹ • (f x' - f x - f' (x' - x))) L (𝓝 0) := -begin - symmetry, - rw [tendsto_iff_norm_tendsto_zero], refine tendsto_congr (λ x', _), - have : ‖x' - x‖⁻¹ ≥ 0, from inv_nonneg.mpr (norm_nonneg _), - simp [norm_smul, abs_of_nonneg this] -end - -lemma has_fderiv_at.lim_real (hf : has_fderiv_at f f' x) (v : E) : - tendsto (λ (c:ℝ), c • (f (x + c⁻¹ • v) - f x)) at_top (𝓝 (f' v)) := -begin - apply hf.lim v, - rw tendsto_at_top_at_top, - exact λ b, ⟨b, λ a ha, le_trans ha (le_abs_self _)⟩ -end - -end - -section tangent_cone - -variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] -{E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] -{F : Type*} [normed_add_comm_group F] [normed_space 𝕜 F] -{f : E → F} {s : set E} {f' : E →L[𝕜] F} - -/-- The image of a tangent cone under the differential of a map is included in the tangent cone to -the image. -/ -lemma has_fderiv_within_at.maps_to_tangent_cone {x : E} (h : has_fderiv_within_at f f' s x) : - maps_to f' (tangent_cone_at 𝕜 s x) (tangent_cone_at 𝕜 (f '' s) (f x)) := -begin - rintros v ⟨c, d, dtop, clim, cdlim⟩, - refine ⟨c, (λn, f (x + d n) - f x), mem_of_superset dtop _, clim, - h.lim at_top dtop clim cdlim⟩, - simp [-mem_image, mem_image_of_mem] {contextual := tt} -end - -/-- If a set has the unique differentiability property at a point x, then the image of this set -under a map with onto derivative has also the unique differentiability property at the image point. --/ -lemma has_fderiv_within_at.unique_diff_within_at {x : E} (h : has_fderiv_within_at f f' s x) - (hs : unique_diff_within_at 𝕜 s x) (h' : dense_range f') : - unique_diff_within_at 𝕜 (f '' s) (f x) := -begin - refine ⟨h'.dense_of_maps_to f'.continuous hs.1 _, - h.continuous_within_at.mem_closure_image hs.2⟩, - show submodule.span 𝕜 (tangent_cone_at 𝕜 s x) ≤ - (submodule.span 𝕜 (tangent_cone_at 𝕜 (f '' s) (f x))).comap f', - rw [submodule.span_le], - exact h.maps_to_tangent_cone.mono (subset.refl _) submodule.subset_span -end - -lemma unique_diff_on.image {f' : E → E →L[𝕜] F} (hs : unique_diff_on 𝕜 s) - (hf' : ∀ x ∈ s, has_fderiv_within_at f (f' x) s x) (hd : ∀ x ∈ s, dense_range (f' x)) : - unique_diff_on 𝕜 (f '' s) := -ball_image_iff.2 $ λ x hx, (hf' x hx).unique_diff_within_at (hs x hx) (hd x hx) - -lemma has_fderiv_within_at.unique_diff_within_at_of_continuous_linear_equiv - {x : E} (e' : E ≃L[𝕜] F) (h : has_fderiv_within_at f (e' : E →L[𝕜] F) s x) - (hs : unique_diff_within_at 𝕜 s x) : - unique_diff_within_at 𝕜 (f '' s) (f x) := -h.unique_diff_within_at hs e'.surjective.dense_range - -lemma continuous_linear_equiv.unique_diff_on_image (e : E ≃L[𝕜] F) (h : unique_diff_on 𝕜 s) : - unique_diff_on 𝕜 (e '' s) := -h.image (λ x _, e.has_fderiv_within_at) (λ x hx, e.surjective.dense_range) - -@[simp] lemma continuous_linear_equiv.unique_diff_on_image_iff (e : E ≃L[𝕜] F) : - unique_diff_on 𝕜 (e '' s) ↔ unique_diff_on 𝕜 s := -⟨λ h, e.symm_image_image s ▸ e.symm.unique_diff_on_image h, e.unique_diff_on_image⟩ - -@[simp] lemma continuous_linear_equiv.unique_diff_on_preimage_iff (e : F ≃L[𝕜] E) : - unique_diff_on 𝕜 (e ⁻¹' s) ↔ unique_diff_on 𝕜 s := -by rw [← e.image_symm_eq_preimage, e.symm.unique_diff_on_image_iff] - -end tangent_cone - -section restrict_scalars -/-! -### Restricting from `ℂ` to `ℝ`, or generally from `𝕜'` to `𝕜` - -If a function is differentiable over `ℂ`, then it is differentiable over `ℝ`. In this paragraph, -we give variants of this statement, in the general situation where `ℂ` and `ℝ` are replaced -respectively by `𝕜'` and `𝕜` where `𝕜'` is a normed algebra over `𝕜`. --/ - -variables (𝕜 : Type*) [nontrivially_normed_field 𝕜] -variables {𝕜' : Type*} [nontrivially_normed_field 𝕜'] [normed_algebra 𝕜 𝕜'] -variables {E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] [normed_space 𝕜' E] -variables [is_scalar_tower 𝕜 𝕜' E] -variables {F : Type*} [normed_add_comm_group F] [normed_space 𝕜 F] [normed_space 𝕜' F] -variables [is_scalar_tower 𝕜 𝕜' F] -variables {f : E → F} {f' : E →L[𝕜'] F} {s : set E} {x : E} - -lemma has_strict_fderiv_at.restrict_scalars (h : has_strict_fderiv_at f f' x) : - has_strict_fderiv_at f (f'.restrict_scalars 𝕜) x := h - -lemma has_fderiv_at_filter.restrict_scalars {L} (h : has_fderiv_at_filter f f' x L) : - has_fderiv_at_filter f (f'.restrict_scalars 𝕜) x L := h - -lemma has_fderiv_at.restrict_scalars (h : has_fderiv_at f f' x) : - has_fderiv_at f (f'.restrict_scalars 𝕜) x := h - -lemma has_fderiv_within_at.restrict_scalars (h : has_fderiv_within_at f f' s x) : - has_fderiv_within_at f (f'.restrict_scalars 𝕜) s x := h - -lemma differentiable_at.restrict_scalars (h : differentiable_at 𝕜' f x) : - differentiable_at 𝕜 f x := -(h.has_fderiv_at.restrict_scalars 𝕜).differentiable_at - -lemma differentiable_within_at.restrict_scalars (h : differentiable_within_at 𝕜' f s x) : - differentiable_within_at 𝕜 f s x := -(h.has_fderiv_within_at.restrict_scalars 𝕜).differentiable_within_at - -lemma differentiable_on.restrict_scalars (h : differentiable_on 𝕜' f s) : - differentiable_on 𝕜 f s := -λx hx, (h x hx).restrict_scalars 𝕜 - -lemma differentiable.restrict_scalars (h : differentiable 𝕜' f) : - differentiable 𝕜 f := -λx, (h x).restrict_scalars 𝕜 - -lemma has_fderiv_within_at_of_restrict_scalars - {g' : E →L[𝕜] F} (h : has_fderiv_within_at f g' s x) - (H : f'.restrict_scalars 𝕜 = g') : has_fderiv_within_at f f' s x := -by { rw ← H at h, exact h } - -lemma has_fderiv_at_of_restrict_scalars {g' : E →L[𝕜] F} (h : has_fderiv_at f g' x) - (H : f'.restrict_scalars 𝕜 = g') : has_fderiv_at f f' x := -by { rw ← H at h, exact h } - -lemma differentiable_at.fderiv_restrict_scalars (h : differentiable_at 𝕜' f x) : - fderiv 𝕜 f x = (fderiv 𝕜' f x).restrict_scalars 𝕜 := -(h.has_fderiv_at.restrict_scalars 𝕜).fderiv - -lemma differentiable_within_at_iff_restrict_scalars - (hf : differentiable_within_at 𝕜 f s x) (hs : unique_diff_within_at 𝕜 s x) : - differentiable_within_at 𝕜' f s x ↔ - ∃ (g' : E →L[𝕜'] F), g'.restrict_scalars 𝕜 = fderiv_within 𝕜 f s x := -begin - split, - { rintros ⟨g', hg'⟩, - exact ⟨g', hs.eq (hg'.restrict_scalars 𝕜) hf.has_fderiv_within_at⟩, }, - { rintros ⟨f', hf'⟩, - exact ⟨f', has_fderiv_within_at_of_restrict_scalars 𝕜 hf.has_fderiv_within_at hf'⟩, }, -end - -lemma differentiable_at_iff_restrict_scalars (hf : differentiable_at 𝕜 f x) : - differentiable_at 𝕜' f x ↔ ∃ (g' : E →L[𝕜'] F), g'.restrict_scalars 𝕜 = fderiv 𝕜 f x := -begin - rw [← differentiable_within_at_univ, ← fderiv_within_univ], - exact differentiable_within_at_iff_restrict_scalars 𝕜 - hf.differentiable_within_at unique_diff_within_at_univ, -end - -end restrict_scalars - /-! ### Support of derivatives -/ section support diff --git a/src/analysis/calculus/fderiv/bilinear.lean b/src/analysis/calculus/fderiv/bilinear.lean new file mode 100644 index 0000000000000..e37f6b4dcadf2 --- /dev/null +++ b/src/analysis/calculus/fderiv/bilinear.lean @@ -0,0 +1,139 @@ +/- +Copyright (c) 2019 Jeremy Avigad. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jeremy Avigad, Sébastien Gouëzel, Yury Kudryashov +-/ +import analysis.calculus.fderiv.prod + +/-! +# The derivative of bounded bilinear maps + +For detailed documentation of the Fréchet derivative, +see the module docstring of `analysis/calculus/fderiv/basic.lean`. + +This file contains the usual formulas (and existence assertions) for the derivative of +bounded bilinear maps. +-/ + +open filter asymptotics continuous_linear_map set metric +open_locale topology classical nnreal filter asymptotics ennreal + +noncomputable theory + + +section + +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] +variables {E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] +variables {F : Type*} [normed_add_comm_group F] [normed_space 𝕜 F] +variables {G : Type*} [normed_add_comm_group G] [normed_space 𝕜 G] +variables {G' : Type*} [normed_add_comm_group G'] [normed_space 𝕜 G'] + +variables {f f₀ f₁ g : E → F} +variables {f' f₀' f₁' g' : E →L[𝕜] F} +variables (e : E →L[𝕜] F) +variables {x : E} +variables {s t : set E} +variables {L L₁ L₂ : filter E} + +section bilinear_map +/-! ### Derivative of a bounded bilinear map -/ + +variables {b : E × F → G} {u : set (E × F)} + +open normed_field + +lemma is_bounded_bilinear_map.has_strict_fderiv_at (h : is_bounded_bilinear_map 𝕜 b) (p : E × F) : + has_strict_fderiv_at b (h.deriv p) p := +begin + rw has_strict_fderiv_at, + set T := (E × F) × (E × F), + have : (λ q : T, b (q.1 - q.2)) =o[𝓝 (p, p)] (λ q : T, ‖q.1 - q.2‖ * 1), + { refine (h.is_O'.comp_tendsto le_top).trans_is_o _, + simp only [(∘)], + refine (is_O_refl (λ q : T, ‖q.1 - q.2‖) _).mul_is_o (is_o.norm_left $ (is_o_one_iff _).2 _), + rw [← sub_self p], + exact continuous_at_fst.sub continuous_at_snd }, + simp only [mul_one, is_o_norm_right] at this, + refine (is_o.congr_of_sub _).1 this, clear this, + convert_to (λ q : T, h.deriv (p - q.2) (q.1 - q.2)) =o[𝓝 (p, p)] (λ q : T, q.1 - q.2), + { ext ⟨⟨x₁, y₁⟩, ⟨x₂, y₂⟩⟩, rcases p with ⟨x, y⟩, + simp only [is_bounded_bilinear_map_deriv_coe, prod.mk_sub_mk, h.map_sub_left, h.map_sub_right], + abel }, + have : (λ q : T, p - q.2) =o[𝓝 (p, p)] (λ q, (1:ℝ)), + from (is_o_one_iff _).2 (sub_self p ▸ tendsto_const_nhds.sub continuous_at_snd), + apply is_bounded_bilinear_map_apply.is_O_comp.trans_is_o, + refine is_o.trans_is_O _ (is_O_const_mul_self 1 _ _).of_norm_right, + refine is_o.mul_is_O _ (is_O_refl _ _), + exact (((h.is_bounded_linear_map_deriv.is_O_id ⊤).comp_tendsto le_top : _).trans_is_o + this).norm_left +end + +lemma is_bounded_bilinear_map.has_fderiv_at (h : is_bounded_bilinear_map 𝕜 b) (p : E × F) : + has_fderiv_at b (h.deriv p) p := +(h.has_strict_fderiv_at p).has_fderiv_at + +lemma is_bounded_bilinear_map.has_fderiv_within_at (h : is_bounded_bilinear_map 𝕜 b) (p : E × F) : + has_fderiv_within_at b (h.deriv p) u p := +(h.has_fderiv_at p).has_fderiv_within_at + +lemma is_bounded_bilinear_map.differentiable_at (h : is_bounded_bilinear_map 𝕜 b) (p : E × F) : + differentiable_at 𝕜 b p := +(h.has_fderiv_at p).differentiable_at + +lemma is_bounded_bilinear_map.differentiable_within_at (h : is_bounded_bilinear_map 𝕜 b) + (p : E × F) : + differentiable_within_at 𝕜 b u p := +(h.differentiable_at p).differentiable_within_at + +lemma is_bounded_bilinear_map.fderiv (h : is_bounded_bilinear_map 𝕜 b) (p : E × F) : + fderiv 𝕜 b p = h.deriv p := +has_fderiv_at.fderiv (h.has_fderiv_at p) + +lemma is_bounded_bilinear_map.fderiv_within (h : is_bounded_bilinear_map 𝕜 b) (p : E × F) + (hxs : unique_diff_within_at 𝕜 u p) : fderiv_within 𝕜 b u p = h.deriv p := +begin + rw differentiable_at.fderiv_within (h.differentiable_at p) hxs, + exact h.fderiv p +end + +lemma is_bounded_bilinear_map.differentiable (h : is_bounded_bilinear_map 𝕜 b) : + differentiable 𝕜 b := +λx, h.differentiable_at x + +lemma is_bounded_bilinear_map.differentiable_on (h : is_bounded_bilinear_map 𝕜 b) : + differentiable_on 𝕜 b u := +h.differentiable.differentiable_on + +variable (B : E →L[𝕜] F →L[𝕜] G) + +lemma continuous_linear_map.has_fderiv_within_at_of_bilinear + {f : G' → E} {g : G' → F} {f' : G' →L[𝕜] E} {g' : G' →L[𝕜] F} {x : G'} {s : set G'} + (hf : has_fderiv_within_at f f' s x) (hg : has_fderiv_within_at g g' s x) : + has_fderiv_within_at (λ y, B (f y) (g y)) (B.precompR G' (f x) g' + B.precompL G' f' (g x)) s x := +(B.is_bounded_bilinear_map.has_fderiv_at (f x, g x)).comp_has_fderiv_within_at x (hf.prod hg) + +lemma continuous_linear_map.has_fderiv_at_of_bilinear + {f : G' → E} {g : G' → F} {f' : G' →L[𝕜] E} {g' : G' →L[𝕜] F} {x : G'} + (hf : has_fderiv_at f f' x) (hg : has_fderiv_at g g' x) : + has_fderiv_at (λ y, B (f y) (g y)) (B.precompR G' (f x) g' + B.precompL G' f' (g x)) x := +(B.is_bounded_bilinear_map.has_fderiv_at (f x, g x)).comp x (hf.prod hg) + +lemma continuous_linear_map.fderiv_within_of_bilinear + {f : G' → E} {g : G' → F} {x : G'} {s : set G'} + (hf : differentiable_within_at 𝕜 f s x) (hg : differentiable_within_at 𝕜 g s x) + (hs : unique_diff_within_at 𝕜 s x) : + fderiv_within 𝕜 (λ y, B (f y) (g y)) s x = + (B.precompR G' (f x) (fderiv_within 𝕜 g s x) + B.precompL G' (fderiv_within 𝕜 f s x) (g x)) := +(B.has_fderiv_within_at_of_bilinear hf.has_fderiv_within_at hg.has_fderiv_within_at).fderiv_within + hs + +lemma continuous_linear_map.fderiv_of_bilinear {f : G' → E} {g : G' → F} {x : G'} + (hf : differentiable_at 𝕜 f x) (hg : differentiable_at 𝕜 g x) : + fderiv 𝕜 (λ y, B (f y) (g y)) x = + (B.precompR G' (f x) (fderiv 𝕜 g x) + B.precompL G' (fderiv 𝕜 f x) (g x)) := +(B.has_fderiv_at_of_bilinear hf.has_fderiv_at hg.has_fderiv_at).fderiv + +end bilinear_map + +end diff --git a/src/analysis/calculus/fderiv/comp.lean b/src/analysis/calculus/fderiv/comp.lean new file mode 100644 index 0000000000000..9304da4ab33cd --- /dev/null +++ b/src/analysis/calculus/fderiv/comp.lean @@ -0,0 +1,243 @@ +/- +Copyright (c) 2019 Jeremy Avigad. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jeremy Avigad, Sébastien Gouëzel, Yury Kudryashov +-/ +import analysis.calculus.fderiv.basic + +/-! +# The derivative of a composition (chain rule) + +For detailed documentation of the Fréchet derivative, +see the module docstring of `analysis/calculus/fderiv/basic.lean`. + +This file contains the usual formulas (and existence assertions) for the derivative of +composition of functions (the chain rule). +-/ + +open filter asymptotics continuous_linear_map set metric +open_locale topology classical nnreal filter asymptotics ennreal + +noncomputable theory + + +section + +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] +variables {E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] +variables {F : Type*} [normed_add_comm_group F] [normed_space 𝕜 F] +variables {G : Type*} [normed_add_comm_group G] [normed_space 𝕜 G] +variables {G' : Type*} [normed_add_comm_group G'] [normed_space 𝕜 G'] + +variables {f f₀ f₁ g : E → F} +variables {f' f₀' f₁' g' : E →L[𝕜] F} +variables (e : E →L[𝕜] F) +variables {x : E} +variables {s t : set E} +variables {L L₁ L₂ : filter E} + + +section composition +/-! +### Derivative of the composition of two functions + +For composition lemmas, we put x explicit to help the elaborator, as otherwise Lean tends to +get confused since there are too many possibilities for composition -/ + +variable (x) + +theorem has_fderiv_at_filter.comp {g : F → G} {g' : F →L[𝕜] G} {L' : filter F} + (hg : has_fderiv_at_filter g g' (f x) L') + (hf : has_fderiv_at_filter f f' x L) (hL : tendsto f L L') : + has_fderiv_at_filter (g ∘ f) (g'.comp f') x L := +let eq₁ := (g'.is_O_comp _ _).trans_is_o hf in +let eq₂ := (hg.comp_tendsto hL).trans_is_O hf.is_O_sub in +by { refine eq₂.triangle (eq₁.congr_left (λ x', _)), simp } + +/- A readable version of the previous theorem, + a general form of the chain rule. -/ + +example {g : F → G} {g' : F →L[𝕜] G} + (hg : has_fderiv_at_filter g g' (f x) (L.map f)) + (hf : has_fderiv_at_filter f f' x L) : + has_fderiv_at_filter (g ∘ f) (g'.comp f') x L := +begin + unfold has_fderiv_at_filter at hg, + have := calc (λ x', g (f x') - g (f x) - g' (f x' - f x)) =o[L] (λ x', f x' - f x) : + hg.comp_tendsto le_rfl + ... =O[L] (λ x', x' - x) : hf.is_O_sub, + refine this.triangle _, + calc (λ x' : E, g' (f x' - f x) - g'.comp f' (x' - x)) + =ᶠ[L] λ x', g' (f x' - f x - f' (x' - x)) : eventually_of_forall (λ x', by simp) + ... =O[L] λ x', f x' - f x - f' (x' - x) : g'.is_O_comp _ _ + ... =o[L] λ x', x' - x : hf +end + +theorem has_fderiv_within_at.comp {g : F → G} {g' : F →L[𝕜] G} {t : set F} + (hg : has_fderiv_within_at g g' t (f x)) (hf : has_fderiv_within_at f f' s x) + (hst : maps_to f s t) : + has_fderiv_within_at (g ∘ f) (g'.comp f') s x := +hg.comp x hf $ hf.continuous_within_at.tendsto_nhds_within hst + +theorem has_fderiv_at.comp_has_fderiv_within_at {g : F → G} {g' : F →L[𝕜] G} + (hg : has_fderiv_at g g' (f x)) (hf : has_fderiv_within_at f f' s x) : + has_fderiv_within_at (g ∘ f) (g'.comp f') s x := +hg.comp x hf hf.continuous_within_at + +theorem has_fderiv_within_at.comp_of_mem {g : F → G} {g' : F →L[𝕜] G} {t : set F} + (hg : has_fderiv_within_at g g' t (f x)) (hf : has_fderiv_within_at f f' s x) + (hst : tendsto f (𝓝[s] x) (𝓝[t] f x)) : + has_fderiv_within_at (g ∘ f) (g'.comp f') s x := +has_fderiv_at_filter.comp x hg hf hst + +/-- The chain rule. -/ +theorem has_fderiv_at.comp {g : F → G} {g' : F →L[𝕜] G} + (hg : has_fderiv_at g g' (f x)) (hf : has_fderiv_at f f' x) : + has_fderiv_at (g ∘ f) (g'.comp f') x := +hg.comp x hf hf.continuous_at + +lemma differentiable_within_at.comp {g : F → G} {t : set F} + (hg : differentiable_within_at 𝕜 g t (f x)) (hf : differentiable_within_at 𝕜 f s x) + (h : maps_to f s t) : differentiable_within_at 𝕜 (g ∘ f) s x := +(hg.has_fderiv_within_at.comp x hf.has_fderiv_within_at h).differentiable_within_at + +lemma differentiable_within_at.comp' {g : F → G} {t : set F} + (hg : differentiable_within_at 𝕜 g t (f x)) (hf : differentiable_within_at 𝕜 f s x) : + differentiable_within_at 𝕜 (g ∘ f) (s ∩ f⁻¹' t) x := +hg.comp x (hf.mono (inter_subset_left _ _)) (inter_subset_right _ _) + +lemma differentiable_at.comp {g : F → G} + (hg : differentiable_at 𝕜 g (f x)) (hf : differentiable_at 𝕜 f x) : + differentiable_at 𝕜 (g ∘ f) x := +(hg.has_fderiv_at.comp x hf.has_fderiv_at).differentiable_at + +lemma differentiable_at.comp_differentiable_within_at {g : F → G} + (hg : differentiable_at 𝕜 g (f x)) (hf : differentiable_within_at 𝕜 f s x) : + differentiable_within_at 𝕜 (g ∘ f) s x := +hg.differentiable_within_at.comp x hf (maps_to_univ _ _) + +lemma fderiv_within.comp {g : F → G} {t : set F} + (hg : differentiable_within_at 𝕜 g t (f x)) (hf : differentiable_within_at 𝕜 f s x) + (h : maps_to f s t) (hxs : unique_diff_within_at 𝕜 s x) : + fderiv_within 𝕜 (g ∘ f) s x = (fderiv_within 𝕜 g t (f x)).comp (fderiv_within 𝕜 f s x) := +(hg.has_fderiv_within_at.comp x (hf.has_fderiv_within_at) h).fderiv_within hxs + +/-- A version of `fderiv_within.comp` that is useful to rewrite the composition of two derivatives + into a single derivative. This version always applies, but creates a new side-goal `f x = y`. -/ +lemma fderiv_within_fderiv_within {g : F → G} {f : E → F} {x : E} {y : F} {s : set E} {t : set F} + (hg : differentiable_within_at 𝕜 g t y) (hf : differentiable_within_at 𝕜 f s x) + (h : maps_to f s t) (hxs : unique_diff_within_at 𝕜 s x) (hy : f x = y) (v : E) : + fderiv_within 𝕜 g t y (fderiv_within 𝕜 f s x v) = fderiv_within 𝕜 (g ∘ f) s x v := +by { subst y, rw [fderiv_within.comp x hg hf h hxs], refl } + +/-- Ternary version of `fderiv_within.comp`, with equality assumptions of basepoints added, in + order to apply more easily as a rewrite from right-to-left. -/ +lemma fderiv_within.comp₃ {g' : G → G'} {g : F → G} {t : set F} {u : set G} {y : F} {y' : G} + (hg' : differentiable_within_at 𝕜 g' u y') (hg : differentiable_within_at 𝕜 g t y) + (hf : differentiable_within_at 𝕜 f s x) + (h2g : maps_to g t u) (h2f : maps_to f s t) + (h3g : g y = y') (h3f : f x = y) (hxs : unique_diff_within_at 𝕜 s x) : + fderiv_within 𝕜 (g' ∘ g ∘ f) s x = (fderiv_within 𝕜 g' u y').comp + ((fderiv_within 𝕜 g t y).comp (fderiv_within 𝕜 f s x)) := +begin + substs h3g h3f, + exact (hg'.has_fderiv_within_at.comp x + (hg.has_fderiv_within_at.comp x (hf.has_fderiv_within_at) h2f) $ h2g.comp h2f).fderiv_within hxs +end + +lemma fderiv.comp {g : F → G} + (hg : differentiable_at 𝕜 g (f x)) (hf : differentiable_at 𝕜 f x) : + fderiv 𝕜 (g ∘ f) x = (fderiv 𝕜 g (f x)).comp (fderiv 𝕜 f x) := +(hg.has_fderiv_at.comp x hf.has_fderiv_at).fderiv + +lemma fderiv.comp_fderiv_within {g : F → G} + (hg : differentiable_at 𝕜 g (f x)) (hf : differentiable_within_at 𝕜 f s x) + (hxs : unique_diff_within_at 𝕜 s x) : + fderiv_within 𝕜 (g ∘ f) s x = (fderiv 𝕜 g (f x)).comp (fderiv_within 𝕜 f s x) := +(hg.has_fderiv_at.comp_has_fderiv_within_at x hf.has_fderiv_within_at).fderiv_within hxs + +lemma differentiable_on.comp {g : F → G} {t : set F} + (hg : differentiable_on 𝕜 g t) (hf : differentiable_on 𝕜 f s) (st : maps_to f s t) : + differentiable_on 𝕜 (g ∘ f) s := +λx hx, differentiable_within_at.comp x (hg (f x) (st hx)) (hf x hx) st + +lemma differentiable.comp {g : F → G} (hg : differentiable 𝕜 g) (hf : differentiable 𝕜 f) : + differentiable 𝕜 (g ∘ f) := +λx, differentiable_at.comp x (hg (f x)) (hf x) + +lemma differentiable.comp_differentiable_on {g : F → G} (hg : differentiable 𝕜 g) + (hf : differentiable_on 𝕜 f s) : + differentiable_on 𝕜 (g ∘ f) s := +hg.differentiable_on.comp hf (maps_to_univ _ _) + +/-- The chain rule for derivatives in the sense of strict differentiability. -/ +protected lemma has_strict_fderiv_at.comp {g : F → G} {g' : F →L[𝕜] G} + (hg : has_strict_fderiv_at g g' (f x)) (hf : has_strict_fderiv_at f f' x) : + has_strict_fderiv_at (λ x, g (f x)) (g'.comp f') x := +((hg.comp_tendsto (hf.continuous_at.prod_map' hf.continuous_at)).trans_is_O hf.is_O_sub).triangle $ + by simpa only [g'.map_sub, f'.coe_comp'] using (g'.is_O_comp _ _).trans_is_o hf + +protected lemma differentiable.iterate {f : E → E} (hf : differentiable 𝕜 f) (n : ℕ) : + differentiable 𝕜 (f^[n]) := +nat.rec_on n differentiable_id (λ n ihn, ihn.comp hf) + +protected lemma differentiable_on.iterate {f : E → E} (hf : differentiable_on 𝕜 f s) + (hs : maps_to f s s) (n : ℕ) : + differentiable_on 𝕜 (f^[n]) s := +nat.rec_on n differentiable_on_id (λ n ihn, ihn.comp hf hs) + +variable {x} + +protected lemma has_fderiv_at_filter.iterate {f : E → E} {f' : E →L[𝕜] E} + (hf : has_fderiv_at_filter f f' x L) (hL : tendsto f L L) (hx : f x = x) (n : ℕ) : + has_fderiv_at_filter (f^[n]) (f'^n) x L := +begin + induction n with n ihn, + { exact has_fderiv_at_filter_id x L }, + { rw [function.iterate_succ, pow_succ'], + rw ← hx at ihn, + exact ihn.comp x hf hL } +end + +protected lemma has_fderiv_at.iterate {f : E → E} {f' : E →L[𝕜] E} + (hf : has_fderiv_at f f' x) (hx : f x = x) (n : ℕ) : + has_fderiv_at (f^[n]) (f'^n) x := +begin + refine hf.iterate _ hx n, + convert hf.continuous_at, + exact hx.symm +end + +protected lemma has_fderiv_within_at.iterate {f : E → E} {f' : E →L[𝕜] E} + (hf : has_fderiv_within_at f f' s x) (hx : f x = x) (hs : maps_to f s s) (n : ℕ) : + has_fderiv_within_at (f^[n]) (f'^n) s x := +begin + refine hf.iterate _ hx n, + convert tendsto_inf.2 ⟨hf.continuous_within_at, _⟩, + exacts [hx.symm, (tendsto_principal_principal.2 hs).mono_left inf_le_right] +end + +protected lemma has_strict_fderiv_at.iterate {f : E → E} {f' : E →L[𝕜] E} + (hf : has_strict_fderiv_at f f' x) (hx : f x = x) (n : ℕ) : + has_strict_fderiv_at (f^[n]) (f'^n) x := +begin + induction n with n ihn, + { exact has_strict_fderiv_at_id x }, + { rw [function.iterate_succ, pow_succ'], + rw ← hx at ihn, + exact ihn.comp x hf } +end + +protected lemma differentiable_at.iterate {f : E → E} (hf : differentiable_at 𝕜 f x) + (hx : f x = x) (n : ℕ) : + differentiable_at 𝕜 (f^[n]) x := +(hf.has_fderiv_at.iterate hx n).differentiable_at + +protected lemma differentiable_within_at.iterate {f : E → E} (hf : differentiable_within_at 𝕜 f s x) + (hx : f x = x) (hs : maps_to f s s) (n : ℕ) : + differentiable_within_at 𝕜 (f^[n]) s x := +(hf.has_fderiv_within_at.iterate hx hs n).differentiable_within_at + +end composition + +end diff --git a/src/analysis/calculus/fderiv/equiv.lean b/src/analysis/calculus/fderiv/equiv.lean new file mode 100644 index 0000000000000..1795a6b6494a0 --- /dev/null +++ b/src/analysis/calculus/fderiv/equiv.lean @@ -0,0 +1,504 @@ +/- +Copyright (c) 2019 Jeremy Avigad. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jeremy Avigad, Sébastien Gouëzel, Yury Kudryashov +-/ +import analysis.calculus.fderiv.linear +import analysis.calculus.fderiv.comp + +/-! +# The derivative of a linear equivalence + +For detailed documentation of the Fréchet derivative, +see the module docstring of `analysis/calculus/fderiv/basic.lean`. + +This file contains the usual formulas (and existence assertions) for the derivative of +continuous linear equivalences. +-/ + +open filter asymptotics continuous_linear_map set metric +open_locale topology classical nnreal filter asymptotics ennreal + +noncomputable theory + + +section + +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] +variables {E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] +variables {F : Type*} [normed_add_comm_group F] [normed_space 𝕜 F] +variables {G : Type*} [normed_add_comm_group G] [normed_space 𝕜 G] +variables {G' : Type*} [normed_add_comm_group G'] [normed_space 𝕜 G'] + +variables {f f₀ f₁ g : E → F} +variables {f' f₀' f₁' g' : E →L[𝕜] F} +variables (e : E →L[𝕜] F) +variables {x : E} +variables {s t : set E} +variables {L L₁ L₂ : filter E} + +namespace continuous_linear_equiv +/-! ### Differentiability of linear equivs, and invariance of differentiability -/ + +variable (iso : E ≃L[𝕜] F) + +protected lemma has_strict_fderiv_at : + has_strict_fderiv_at iso (iso : E →L[𝕜] F) x := +iso.to_continuous_linear_map.has_strict_fderiv_at + +protected lemma has_fderiv_within_at : + has_fderiv_within_at iso (iso : E →L[𝕜] F) s x := +iso.to_continuous_linear_map.has_fderiv_within_at + +protected lemma has_fderiv_at : has_fderiv_at iso (iso : E →L[𝕜] F) x := +iso.to_continuous_linear_map.has_fderiv_at_filter + +protected lemma differentiable_at : differentiable_at 𝕜 iso x := +iso.has_fderiv_at.differentiable_at + +protected lemma differentiable_within_at : + differentiable_within_at 𝕜 iso s x := +iso.differentiable_at.differentiable_within_at + +protected lemma fderiv : fderiv 𝕜 iso x = iso := +iso.has_fderiv_at.fderiv + +protected lemma fderiv_within (hxs : unique_diff_within_at 𝕜 s x) : + fderiv_within 𝕜 iso s x = iso := +iso.to_continuous_linear_map.fderiv_within hxs + +protected lemma differentiable : differentiable 𝕜 iso := +λx, iso.differentiable_at + +protected lemma differentiable_on : differentiable_on 𝕜 iso s := +iso.differentiable.differentiable_on + +lemma comp_differentiable_within_at_iff {f : G → E} {s : set G} {x : G} : + differentiable_within_at 𝕜 (iso ∘ f) s x ↔ differentiable_within_at 𝕜 f s x := +begin + refine ⟨λ H, _, λ H, iso.differentiable.differentiable_at.comp_differentiable_within_at x H⟩, + have : differentiable_within_at 𝕜 (iso.symm ∘ (iso ∘ f)) s x := + iso.symm.differentiable.differentiable_at.comp_differentiable_within_at x H, + rwa [← function.comp.assoc iso.symm iso f, iso.symm_comp_self] at this, +end + +lemma comp_differentiable_at_iff {f : G → E} {x : G} : + differentiable_at 𝕜 (iso ∘ f) x ↔ differentiable_at 𝕜 f x := +by rw [← differentiable_within_at_univ, ← differentiable_within_at_univ, + iso.comp_differentiable_within_at_iff] + +lemma comp_differentiable_on_iff {f : G → E} {s : set G} : + differentiable_on 𝕜 (iso ∘ f) s ↔ differentiable_on 𝕜 f s := +begin + rw [differentiable_on, differentiable_on], + simp only [iso.comp_differentiable_within_at_iff], +end + +lemma comp_differentiable_iff {f : G → E} : + differentiable 𝕜 (iso ∘ f) ↔ differentiable 𝕜 f := +begin + rw [← differentiable_on_univ, ← differentiable_on_univ], + exact iso.comp_differentiable_on_iff +end + +lemma comp_has_fderiv_within_at_iff + {f : G → E} {s : set G} {x : G} {f' : G →L[𝕜] E} : + has_fderiv_within_at (iso ∘ f) ((iso : E →L[𝕜] F).comp f') s x ↔ has_fderiv_within_at f f' s x := +begin + refine ⟨λ H, _, λ H, iso.has_fderiv_at.comp_has_fderiv_within_at x H⟩, + have A : f = iso.symm ∘ (iso ∘ f), by { rw [← function.comp.assoc, iso.symm_comp_self], refl }, + have B : f' = (iso.symm : F →L[𝕜] E).comp ((iso : E →L[𝕜] F).comp f'), + by rw [← continuous_linear_map.comp_assoc, iso.coe_symm_comp_coe, + continuous_linear_map.id_comp], + rw [A, B], + exact iso.symm.has_fderiv_at.comp_has_fderiv_within_at x H +end + +lemma comp_has_strict_fderiv_at_iff {f : G → E} {x : G} {f' : G →L[𝕜] E} : + has_strict_fderiv_at (iso ∘ f) ((iso : E →L[𝕜] F).comp f') x ↔ has_strict_fderiv_at f f' x := +begin + refine ⟨λ H, _, λ H, iso.has_strict_fderiv_at.comp x H⟩, + convert iso.symm.has_strict_fderiv_at.comp x H; ext z; apply (iso.symm_apply_apply _).symm +end + +lemma comp_has_fderiv_at_iff {f : G → E} {x : G} {f' : G →L[𝕜] E} : + has_fderiv_at (iso ∘ f) ((iso : E →L[𝕜] F).comp f') x ↔ has_fderiv_at f f' x := +by simp_rw [← has_fderiv_within_at_univ, iso.comp_has_fderiv_within_at_iff] + +lemma comp_has_fderiv_within_at_iff' + {f : G → E} {s : set G} {x : G} {f' : G →L[𝕜] F} : + has_fderiv_within_at (iso ∘ f) f' s x ↔ + has_fderiv_within_at f ((iso.symm : F →L[𝕜] E).comp f') s x := +by rw [← iso.comp_has_fderiv_within_at_iff, ← continuous_linear_map.comp_assoc, + iso.coe_comp_coe_symm, continuous_linear_map.id_comp] + +lemma comp_has_fderiv_at_iff' {f : G → E} {x : G} {f' : G →L[𝕜] F} : + has_fderiv_at (iso ∘ f) f' x ↔ has_fderiv_at f ((iso.symm : F →L[𝕜] E).comp f') x := +by simp_rw [← has_fderiv_within_at_univ, iso.comp_has_fderiv_within_at_iff'] + +lemma comp_fderiv_within {f : G → E} {s : set G} {x : G} + (hxs : unique_diff_within_at 𝕜 s x) : + fderiv_within 𝕜 (iso ∘ f) s x = (iso : E →L[𝕜] F).comp (fderiv_within 𝕜 f s x) := +begin + by_cases h : differentiable_within_at 𝕜 f s x, + { rw [fderiv.comp_fderiv_within x iso.differentiable_at h hxs, iso.fderiv] }, + { have : ¬differentiable_within_at 𝕜 (iso ∘ f) s x, + from mt iso.comp_differentiable_within_at_iff.1 h, + rw [fderiv_within_zero_of_not_differentiable_within_at h, + fderiv_within_zero_of_not_differentiable_within_at this, + continuous_linear_map.comp_zero] } +end + +lemma comp_fderiv {f : G → E} {x : G} : + fderiv 𝕜 (iso ∘ f) x = (iso : E →L[𝕜] F).comp (fderiv 𝕜 f x) := +begin + rw [← fderiv_within_univ, ← fderiv_within_univ], + exact iso.comp_fderiv_within unique_diff_within_at_univ, +end + +lemma comp_right_differentiable_within_at_iff {f : F → G} {s : set F} {x : E} : + differentiable_within_at 𝕜 (f ∘ iso) (iso ⁻¹' s) x ↔ differentiable_within_at 𝕜 f s (iso x) := +begin + refine ⟨λ H, _, λ H, H.comp x iso.differentiable_within_at (maps_to_preimage _ s)⟩, + have : differentiable_within_at 𝕜 ((f ∘ iso) ∘ iso.symm) s (iso x), + { rw ← iso.symm_apply_apply x at H, + apply H.comp (iso x) iso.symm.differentiable_within_at, + assume y hy, + simpa only [mem_preimage, apply_symm_apply] using hy }, + rwa [function.comp.assoc, iso.self_comp_symm] at this, +end + +lemma comp_right_differentiable_at_iff {f : F → G} {x : E} : + differentiable_at 𝕜 (f ∘ iso) x ↔ differentiable_at 𝕜 f (iso x) := +by simp only [← differentiable_within_at_univ, ← iso.comp_right_differentiable_within_at_iff, + preimage_univ] + +lemma comp_right_differentiable_on_iff {f : F → G} {s : set F} : + differentiable_on 𝕜 (f ∘ iso) (iso ⁻¹' s) ↔ differentiable_on 𝕜 f s := +begin + refine ⟨λ H y hy, _, λ H y hy, iso.comp_right_differentiable_within_at_iff.2 (H _ hy)⟩, + rw [← iso.apply_symm_apply y, ← comp_right_differentiable_within_at_iff], + apply H, + simpa only [mem_preimage, apply_symm_apply] using hy, +end + +lemma comp_right_differentiable_iff {f : F → G} : + differentiable 𝕜 (f ∘ iso) ↔ differentiable 𝕜 f := +by simp only [← differentiable_on_univ, ← iso.comp_right_differentiable_on_iff, preimage_univ] + +lemma comp_right_has_fderiv_within_at_iff + {f : F → G} {s : set F} {x : E} {f' : F →L[𝕜] G} : + has_fderiv_within_at (f ∘ iso) (f'.comp (iso : E →L[𝕜] F)) (iso ⁻¹' s) x ↔ + has_fderiv_within_at f f' s (iso x) := +begin + refine ⟨λ H, _, λ H, H.comp x iso.has_fderiv_within_at (maps_to_preimage _ s)⟩, + rw [← iso.symm_apply_apply x] at H, + have A : f = (f ∘ iso) ∘ iso.symm, by { rw [function.comp.assoc, iso.self_comp_symm], refl }, + have B : f' = (f'.comp (iso : E →L[𝕜] F)).comp (iso.symm : F →L[𝕜] E), + by rw [continuous_linear_map.comp_assoc, iso.coe_comp_coe_symm, + continuous_linear_map.comp_id], + rw [A, B], + apply H.comp (iso x) iso.symm.has_fderiv_within_at, + assume y hy, + simpa only [mem_preimage, apply_symm_apply] using hy +end + +lemma comp_right_has_fderiv_at_iff {f : F → G} {x : E} {f' : F →L[𝕜] G} : + has_fderiv_at (f ∘ iso) (f'.comp (iso : E →L[𝕜] F)) x ↔ has_fderiv_at f f' (iso x) := +by simp only [← has_fderiv_within_at_univ, ← comp_right_has_fderiv_within_at_iff, preimage_univ] + +lemma comp_right_has_fderiv_within_at_iff' + {f : F → G} {s : set F} {x : E} {f' : E →L[𝕜] G} : + has_fderiv_within_at (f ∘ iso) f' (iso ⁻¹' s) x ↔ + has_fderiv_within_at f (f'.comp (iso.symm : F →L[𝕜] E)) s (iso x) := +by rw [← iso.comp_right_has_fderiv_within_at_iff, continuous_linear_map.comp_assoc, + iso.coe_symm_comp_coe, continuous_linear_map.comp_id] + +lemma comp_right_has_fderiv_at_iff' {f : F → G} {x : E} {f' : E →L[𝕜] G} : + has_fderiv_at (f ∘ iso) f' x ↔ has_fderiv_at f (f'.comp (iso.symm : F →L[𝕜] E)) (iso x) := +by simp only [← has_fderiv_within_at_univ, ← iso.comp_right_has_fderiv_within_at_iff', + preimage_univ] + +lemma comp_right_fderiv_within {f : F → G} {s : set F} {x : E} + (hxs : unique_diff_within_at 𝕜 (iso ⁻¹' s) x) : + fderiv_within 𝕜 (f ∘ iso) (iso ⁻¹'s) x = (fderiv_within 𝕜 f s (iso x)).comp (iso : E →L[𝕜] F) := +begin + by_cases h : differentiable_within_at 𝕜 f s (iso x), + { exact (iso.comp_right_has_fderiv_within_at_iff.2 (h.has_fderiv_within_at)).fderiv_within hxs }, + { have : ¬ differentiable_within_at 𝕜 (f ∘ iso) (iso ⁻¹' s) x, + { assume h', exact h (iso.comp_right_differentiable_within_at_iff.1 h') }, + rw [fderiv_within_zero_of_not_differentiable_within_at h, + fderiv_within_zero_of_not_differentiable_within_at this, continuous_linear_map.zero_comp] } +end + +lemma comp_right_fderiv {f : F → G} {x : E} : + fderiv 𝕜 (f ∘ iso) x = (fderiv 𝕜 f (iso x)).comp (iso : E →L[𝕜] F) := +begin + rw [← fderiv_within_univ, ← fderiv_within_univ, ← iso.comp_right_fderiv_within, preimage_univ], + exact unique_diff_within_at_univ, +end + +end continuous_linear_equiv + +namespace linear_isometry_equiv +/-! ### Differentiability of linear isometry equivs, and invariance of differentiability -/ + +variable (iso : E ≃ₗᵢ[𝕜] F) + +protected lemma has_strict_fderiv_at : has_strict_fderiv_at iso (iso : E →L[𝕜] F) x := +(iso : E ≃L[𝕜] F).has_strict_fderiv_at + +protected lemma has_fderiv_within_at : has_fderiv_within_at iso (iso : E →L[𝕜] F) s x := +(iso : E ≃L[𝕜] F).has_fderiv_within_at + +protected lemma has_fderiv_at : has_fderiv_at iso (iso : E →L[𝕜] F) x := +(iso : E ≃L[𝕜] F).has_fderiv_at + +protected lemma differentiable_at : differentiable_at 𝕜 iso x := +iso.has_fderiv_at.differentiable_at + +protected lemma differentiable_within_at : + differentiable_within_at 𝕜 iso s x := +iso.differentiable_at.differentiable_within_at + +protected lemma fderiv : fderiv 𝕜 iso x = iso := iso.has_fderiv_at.fderiv + +protected lemma fderiv_within (hxs : unique_diff_within_at 𝕜 s x) : + fderiv_within 𝕜 iso s x = iso := +(iso : E ≃L[𝕜] F).fderiv_within hxs + +protected lemma differentiable : differentiable 𝕜 iso := +λx, iso.differentiable_at + +protected lemma differentiable_on : differentiable_on 𝕜 iso s := +iso.differentiable.differentiable_on + +lemma comp_differentiable_within_at_iff {f : G → E} {s : set G} {x : G} : + differentiable_within_at 𝕜 (iso ∘ f) s x ↔ differentiable_within_at 𝕜 f s x := +(iso : E ≃L[𝕜] F).comp_differentiable_within_at_iff + +lemma comp_differentiable_at_iff {f : G → E} {x : G} : + differentiable_at 𝕜 (iso ∘ f) x ↔ differentiable_at 𝕜 f x := +(iso : E ≃L[𝕜] F).comp_differentiable_at_iff + +lemma comp_differentiable_on_iff {f : G → E} {s : set G} : + differentiable_on 𝕜 (iso ∘ f) s ↔ differentiable_on 𝕜 f s := +(iso : E ≃L[𝕜] F).comp_differentiable_on_iff + +lemma comp_differentiable_iff {f : G → E} : + differentiable 𝕜 (iso ∘ f) ↔ differentiable 𝕜 f := +(iso : E ≃L[𝕜] F).comp_differentiable_iff + +lemma comp_has_fderiv_within_at_iff + {f : G → E} {s : set G} {x : G} {f' : G →L[𝕜] E} : + has_fderiv_within_at (iso ∘ f) ((iso : E →L[𝕜] F).comp f') s x ↔ has_fderiv_within_at f f' s x := +(iso : E ≃L[𝕜] F).comp_has_fderiv_within_at_iff + +lemma comp_has_strict_fderiv_at_iff {f : G → E} {x : G} {f' : G →L[𝕜] E} : + has_strict_fderiv_at (iso ∘ f) ((iso : E →L[𝕜] F).comp f') x ↔ has_strict_fderiv_at f f' x := +(iso : E ≃L[𝕜] F).comp_has_strict_fderiv_at_iff + +lemma comp_has_fderiv_at_iff {f : G → E} {x : G} {f' : G →L[𝕜] E} : + has_fderiv_at (iso ∘ f) ((iso : E →L[𝕜] F).comp f') x ↔ has_fderiv_at f f' x := +(iso : E ≃L[𝕜] F).comp_has_fderiv_at_iff + +lemma comp_has_fderiv_within_at_iff' + {f : G → E} {s : set G} {x : G} {f' : G →L[𝕜] F} : + has_fderiv_within_at (iso ∘ f) f' s x ↔ + has_fderiv_within_at f ((iso.symm : F →L[𝕜] E).comp f') s x := +(iso : E ≃L[𝕜] F).comp_has_fderiv_within_at_iff' + +lemma comp_has_fderiv_at_iff' {f : G → E} {x : G} {f' : G →L[𝕜] F} : + has_fderiv_at (iso ∘ f) f' x ↔ has_fderiv_at f ((iso.symm : F →L[𝕜] E).comp f') x := +(iso : E ≃L[𝕜] F).comp_has_fderiv_at_iff' + +lemma comp_fderiv_within {f : G → E} {s : set G} {x : G} + (hxs : unique_diff_within_at 𝕜 s x) : + fderiv_within 𝕜 (iso ∘ f) s x = (iso : E →L[𝕜] F).comp (fderiv_within 𝕜 f s x) := +(iso : E ≃L[𝕜] F).comp_fderiv_within hxs + +lemma comp_fderiv {f : G → E} {x : G} : + fderiv 𝕜 (iso ∘ f) x = (iso : E →L[𝕜] F).comp (fderiv 𝕜 f x) := +(iso : E ≃L[𝕜] F).comp_fderiv + +end linear_isometry_equiv + +/-- If `f (g y) = y` for `y` in some neighborhood of `a`, `g` is continuous at `a`, and `f` has an +invertible derivative `f'` at `g a` in the strict sense, then `g` has the derivative `f'⁻¹` at `a` +in the strict sense. + +This is one of the easy parts of the inverse function theorem: it assumes that we already have an +inverse function. -/ +theorem has_strict_fderiv_at.of_local_left_inverse {f : E → F} {f' : E ≃L[𝕜] F} {g : F → E} {a : F} + (hg : continuous_at g a) (hf : has_strict_fderiv_at f (f' : E →L[𝕜] F) (g a)) + (hfg : ∀ᶠ y in 𝓝 a, f (g y) = y) : + has_strict_fderiv_at g (f'.symm : F →L[𝕜] E) a := +begin + replace hg := hg.prod_map' hg, + replace hfg := hfg.prod_mk_nhds hfg, + have : (λ p : F × F, g p.1 - g p.2 - f'.symm (p.1 - p.2)) =O[𝓝 (a, a)] + (λ p : F × F, f' (g p.1 - g p.2) - (p.1 - p.2)), + { refine ((f'.symm : F →L[𝕜] E).is_O_comp _ _).congr (λ x, _) (λ _, rfl), + simp }, + refine this.trans_is_o _, clear this, + refine ((hf.comp_tendsto hg).symm.congr' (hfg.mono _) + (eventually_of_forall $ λ _, rfl)).trans_is_O _, + { rintros p ⟨hp1, hp2⟩, + simp [hp1, hp2] }, + { refine (hf.is_O_sub_rev.comp_tendsto hg).congr' + (eventually_of_forall $ λ _, rfl) (hfg.mono _), + rintros p ⟨hp1, hp2⟩, + simp only [(∘), hp1, hp2] } +end + +/-- If `f (g y) = y` for `y` in some neighborhood of `a`, `g` is continuous at `a`, and `f` has an +invertible derivative `f'` at `g a`, then `g` has the derivative `f'⁻¹` at `a`. + +This is one of the easy parts of the inverse function theorem: it assumes that we already have +an inverse function. -/ +theorem has_fderiv_at.of_local_left_inverse {f : E → F} {f' : E ≃L[𝕜] F} {g : F → E} {a : F} + (hg : continuous_at g a) (hf : has_fderiv_at f (f' : E →L[𝕜] F) (g a)) + (hfg : ∀ᶠ y in 𝓝 a, f (g y) = y) : + has_fderiv_at g (f'.symm : F →L[𝕜] E) a := +begin + have : (λ x : F, g x - g a - f'.symm (x - a)) =O[𝓝 a] (λ x : F, f' (g x - g a) - (x - a)), + { refine ((f'.symm : F →L[𝕜] E).is_O_comp _ _).congr (λ x, _) (λ _, rfl), + simp }, + refine this.trans_is_o _, clear this, + refine ((hf.comp_tendsto hg).symm.congr' (hfg.mono _) + (eventually_of_forall $ λ _, rfl)).trans_is_O _, + { rintros p hp, + simp [hp, hfg.self_of_nhds] }, + { refine ((hf.is_O_sub_rev f'.antilipschitz).comp_tendsto hg).congr' + (eventually_of_forall $ λ _, rfl) (hfg.mono _), + rintros p hp, + simp only [(∘), hp, hfg.self_of_nhds] } +end + +/-- If `f` is a local homeomorphism defined on a neighbourhood of `f.symm a`, and `f` has an +invertible derivative `f'` in the sense of strict differentiability at `f.symm a`, then `f.symm` has +the derivative `f'⁻¹` at `a`. + +This is one of the easy parts of the inverse function theorem: it assumes that we already have +an inverse function. -/ +lemma local_homeomorph.has_strict_fderiv_at_symm (f : local_homeomorph E F) {f' : E ≃L[𝕜] F} {a : F} + (ha : a ∈ f.target) (htff' : has_strict_fderiv_at f (f' : E →L[𝕜] F) (f.symm a)) : + has_strict_fderiv_at f.symm (f'.symm : F →L[𝕜] E) a := +htff'.of_local_left_inverse (f.symm.continuous_at ha) (f.eventually_right_inverse ha) + +/-- If `f` is a local homeomorphism defined on a neighbourhood of `f.symm a`, and `f` has an +invertible derivative `f'` at `f.symm a`, then `f.symm` has the derivative `f'⁻¹` at `a`. + +This is one of the easy parts of the inverse function theorem: it assumes that we already have +an inverse function. -/ +lemma local_homeomorph.has_fderiv_at_symm (f : local_homeomorph E F) {f' : E ≃L[𝕜] F} {a : F} + (ha : a ∈ f.target) (htff' : has_fderiv_at f (f' : E →L[𝕜] F) (f.symm a)) : + has_fderiv_at f.symm (f'.symm : F →L[𝕜] E) a := +htff'.of_local_left_inverse (f.symm.continuous_at ha) (f.eventually_right_inverse ha) + +lemma has_fderiv_within_at.eventually_ne (h : has_fderiv_within_at f f' s x) + (hf' : ∃ C, ∀ z, ‖z‖ ≤ C * ‖f' z‖) : + ∀ᶠ z in 𝓝[s \ {x}] x, f z ≠ f x := +begin + rw [nhds_within, diff_eq, ← inf_principal, ← inf_assoc, eventually_inf_principal], + have A : (λ z, z - x) =O[𝓝[s] x] (λ z, f' (z - x)) := + (is_O_iff.2 $ hf'.imp $ λ C hC, eventually_of_forall $ λ z, hC _), + have : (λ z, f z - f x) ~[𝓝[s] x] (λ z, f' (z - x)) := h.trans_is_O A, + simpa [not_imp_not, sub_eq_zero] using (A.trans this.is_O_symm).eq_zero_imp +end + +lemma has_fderiv_at.eventually_ne (h : has_fderiv_at f f' x) (hf' : ∃ C, ∀ z, ‖z‖ ≤ C * ‖f' z‖) : + ∀ᶠ z in 𝓝[≠] x, f z ≠ f x := +by simpa only [compl_eq_univ_diff] using (has_fderiv_within_at_univ.2 h).eventually_ne hf' + +end + +section +/- + In the special case of a normed space over the reals, + we can use scalar multiplication in the `tendsto` characterization + of the Fréchet derivative. +-/ + + +variables {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] +variables {F : Type*} [normed_add_comm_group F] [normed_space ℝ F] +variables {f : E → F} {f' : E →L[ℝ] F} {x : E} + +theorem has_fderiv_at_filter_real_equiv {L : filter E} : + tendsto (λ x' : E, ‖x' - x‖⁻¹ * ‖f x' - f x - f' (x' - x)‖) L (𝓝 0) ↔ + tendsto (λ x' : E, ‖x' - x‖⁻¹ • (f x' - f x - f' (x' - x))) L (𝓝 0) := +begin + symmetry, + rw [tendsto_iff_norm_tendsto_zero], refine tendsto_congr (λ x', _), + have : ‖x' - x‖⁻¹ ≥ 0, from inv_nonneg.mpr (norm_nonneg _), + simp [norm_smul, abs_of_nonneg this] +end + +lemma has_fderiv_at.lim_real (hf : has_fderiv_at f f' x) (v : E) : + tendsto (λ (c:ℝ), c • (f (x + c⁻¹ • v) - f x)) at_top (𝓝 (f' v)) := +begin + apply hf.lim v, + rw tendsto_at_top_at_top, + exact λ b, ⟨b, λ a ha, le_trans ha (le_abs_self _)⟩ +end + +end + +section tangent_cone + +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] +{E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] +{F : Type*} [normed_add_comm_group F] [normed_space 𝕜 F] +{f : E → F} {s : set E} {f' : E →L[𝕜] F} + +/-- The image of a tangent cone under the differential of a map is included in the tangent cone to +the image. -/ +lemma has_fderiv_within_at.maps_to_tangent_cone {x : E} (h : has_fderiv_within_at f f' s x) : + maps_to f' (tangent_cone_at 𝕜 s x) (tangent_cone_at 𝕜 (f '' s) (f x)) := +begin + rintros v ⟨c, d, dtop, clim, cdlim⟩, + refine ⟨c, (λn, f (x + d n) - f x), mem_of_superset dtop _, clim, + h.lim at_top dtop clim cdlim⟩, + simp [-mem_image, mem_image_of_mem] {contextual := tt} +end + +/-- If a set has the unique differentiability property at a point x, then the image of this set +under a map with onto derivative has also the unique differentiability property at the image point. +-/ +lemma has_fderiv_within_at.unique_diff_within_at {x : E} (h : has_fderiv_within_at f f' s x) + (hs : unique_diff_within_at 𝕜 s x) (h' : dense_range f') : + unique_diff_within_at 𝕜 (f '' s) (f x) := +begin + refine ⟨h'.dense_of_maps_to f'.continuous hs.1 _, + h.continuous_within_at.mem_closure_image hs.2⟩, + show submodule.span 𝕜 (tangent_cone_at 𝕜 s x) ≤ + (submodule.span 𝕜 (tangent_cone_at 𝕜 (f '' s) (f x))).comap f', + rw [submodule.span_le], + exact h.maps_to_tangent_cone.mono (subset.refl _) submodule.subset_span +end + +lemma unique_diff_on.image {f' : E → E →L[𝕜] F} (hs : unique_diff_on 𝕜 s) + (hf' : ∀ x ∈ s, has_fderiv_within_at f (f' x) s x) (hd : ∀ x ∈ s, dense_range (f' x)) : + unique_diff_on 𝕜 (f '' s) := +ball_image_iff.2 $ λ x hx, (hf' x hx).unique_diff_within_at (hs x hx) (hd x hx) + +lemma has_fderiv_within_at.unique_diff_within_at_of_continuous_linear_equiv + {x : E} (e' : E ≃L[𝕜] F) (h : has_fderiv_within_at f (e' : E →L[𝕜] F) s x) + (hs : unique_diff_within_at 𝕜 s x) : + unique_diff_within_at 𝕜 (f '' s) (f x) := +h.unique_diff_within_at hs e'.surjective.dense_range + +lemma continuous_linear_equiv.unique_diff_on_image (e : E ≃L[𝕜] F) (h : unique_diff_on 𝕜 s) : + unique_diff_on 𝕜 (e '' s) := +h.image (λ x _, e.has_fderiv_within_at) (λ x hx, e.surjective.dense_range) + +@[simp] lemma continuous_linear_equiv.unique_diff_on_image_iff (e : E ≃L[𝕜] F) : + unique_diff_on 𝕜 (e '' s) ↔ unique_diff_on 𝕜 s := +⟨λ h, e.symm_image_image s ▸ e.symm.unique_diff_on_image h, e.unique_diff_on_image⟩ + +@[simp] lemma continuous_linear_equiv.unique_diff_on_preimage_iff (e : F ≃L[𝕜] E) : + unique_diff_on 𝕜 (e ⁻¹' s) ↔ unique_diff_on 𝕜 s := +by rw [← e.image_symm_eq_preimage, e.symm.unique_diff_on_image_iff] + +end tangent_cone diff --git a/src/analysis/calculus/fderiv/linear.lean b/src/analysis/calculus/fderiv/linear.lean new file mode 100644 index 0000000000000..28b63c073871e --- /dev/null +++ b/src/analysis/calculus/fderiv/linear.lean @@ -0,0 +1,124 @@ +/- +Copyright (c) 2019 Jeremy Avigad. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jeremy Avigad, Sébastien Gouëzel, Yury Kudryashov +-/ +import analysis.calculus.fderiv.basic + +/-! +# The derivative of bounded linear maps + +For detailed documentation of the Fréchet derivative, +see the module docstring of `analysis/calculus/fderiv/basic.lean`. + +This file contains the usual formulas (and existence assertions) for the derivative of +bounded linear maps. +-/ + +open filter asymptotics continuous_linear_map set metric +open_locale topology classical nnreal filter asymptotics ennreal + +noncomputable theory + + +section + +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] +variables {E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] +variables {F : Type*} [normed_add_comm_group F] [normed_space 𝕜 F] +variables {G : Type*} [normed_add_comm_group G] [normed_space 𝕜 G] +variables {G' : Type*} [normed_add_comm_group G'] [normed_space 𝕜 G'] + +variables {f f₀ f₁ g : E → F} +variables {f' f₀' f₁' g' : E →L[𝕜] F} +variables (e : E →L[𝕜] F) +variables {x : E} +variables {s t : set E} +variables {L L₁ L₂ : filter E} + +section continuous_linear_map +/-! +### Continuous linear maps + +There are currently two variants of these in mathlib, the bundled version +(named `continuous_linear_map`, and denoted `E →L[𝕜] F`), and the unbundled version (with a +predicate `is_bounded_linear_map`). We give statements for both versions. -/ + +protected theorem continuous_linear_map.has_strict_fderiv_at {x : E} : + has_strict_fderiv_at e e x := +(is_o_zero _ _).congr_left $ λ x, by simp only [e.map_sub, sub_self] + +protected lemma continuous_linear_map.has_fderiv_at_filter : + has_fderiv_at_filter e e x L := +(is_o_zero _ _).congr_left $ λ x, by simp only [e.map_sub, sub_self] + +protected lemma continuous_linear_map.has_fderiv_within_at : has_fderiv_within_at e e s x := +e.has_fderiv_at_filter + +protected lemma continuous_linear_map.has_fderiv_at : has_fderiv_at e e x := +e.has_fderiv_at_filter + +@[simp] protected lemma continuous_linear_map.differentiable_at : differentiable_at 𝕜 e x := +e.has_fderiv_at.differentiable_at + +protected lemma continuous_linear_map.differentiable_within_at : differentiable_within_at 𝕜 e s x := +e.differentiable_at.differentiable_within_at + +@[simp] protected lemma continuous_linear_map.fderiv : fderiv 𝕜 e x = e := +e.has_fderiv_at.fderiv + +protected lemma continuous_linear_map.fderiv_within (hxs : unique_diff_within_at 𝕜 s x) : + fderiv_within 𝕜 e s x = e := +begin + rw differentiable_at.fderiv_within e.differentiable_at hxs, + exact e.fderiv +end + +@[simp] protected lemma continuous_linear_map.differentiable : differentiable 𝕜 e := +λx, e.differentiable_at + +protected lemma continuous_linear_map.differentiable_on : differentiable_on 𝕜 e s := +e.differentiable.differentiable_on + +lemma is_bounded_linear_map.has_fderiv_at_filter (h : is_bounded_linear_map 𝕜 f) : + has_fderiv_at_filter f h.to_continuous_linear_map x L := +h.to_continuous_linear_map.has_fderiv_at_filter + +lemma is_bounded_linear_map.has_fderiv_within_at (h : is_bounded_linear_map 𝕜 f) : + has_fderiv_within_at f h.to_continuous_linear_map s x := +h.has_fderiv_at_filter + +lemma is_bounded_linear_map.has_fderiv_at (h : is_bounded_linear_map 𝕜 f) : + has_fderiv_at f h.to_continuous_linear_map x := +h.has_fderiv_at_filter + +lemma is_bounded_linear_map.differentiable_at (h : is_bounded_linear_map 𝕜 f) : + differentiable_at 𝕜 f x := +h.has_fderiv_at.differentiable_at + +lemma is_bounded_linear_map.differentiable_within_at (h : is_bounded_linear_map 𝕜 f) : + differentiable_within_at 𝕜 f s x := +h.differentiable_at.differentiable_within_at + +lemma is_bounded_linear_map.fderiv (h : is_bounded_linear_map 𝕜 f) : + fderiv 𝕜 f x = h.to_continuous_linear_map := +has_fderiv_at.fderiv (h.has_fderiv_at) + +lemma is_bounded_linear_map.fderiv_within (h : is_bounded_linear_map 𝕜 f) + (hxs : unique_diff_within_at 𝕜 s x) : fderiv_within 𝕜 f s x = h.to_continuous_linear_map := +begin + rw differentiable_at.fderiv_within h.differentiable_at hxs, + exact h.fderiv +end + +lemma is_bounded_linear_map.differentiable (h : is_bounded_linear_map 𝕜 f) : + differentiable 𝕜 f := +λx, h.differentiable_at + +lemma is_bounded_linear_map.differentiable_on (h : is_bounded_linear_map 𝕜 f) : + differentiable_on 𝕜 f s := +h.differentiable.differentiable_on + +end continuous_linear_map + +end diff --git a/src/analysis/calculus/fderiv/mul.lean b/src/analysis/calculus/fderiv/mul.lean new file mode 100644 index 0000000000000..203ae912cf5b0 --- /dev/null +++ b/src/analysis/calculus/fderiv/mul.lean @@ -0,0 +1,465 @@ +/- +Copyright (c) 2019 Jeremy Avigad. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jeremy Avigad, Sébastien Gouëzel, Yury Kudryashov +-/ +import analysis.calculus.fderiv.bilinear + +/-! +# Multiplicative operations on derivatives + +For detailed documentation of the Fréchet derivative, +see the module docstring of `analysis/calculus/fderiv/basic.lean`. + +This file contains the usual formulas (and existence assertions) for the derivative of + +* multiplication of a function by a scalar function +* multiplication of two scalar functions +* inverse function (assuming that it exists; the inverse function theorem is in `../inverse.lean`) +-/ + +open filter asymptotics continuous_linear_map set metric +open_locale topology classical nnreal filter asymptotics ennreal + +noncomputable theory + + +section + +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] +variables {E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] +variables {F : Type*} [normed_add_comm_group F] [normed_space 𝕜 F] +variables {G : Type*} [normed_add_comm_group G] [normed_space 𝕜 G] +variables {G' : Type*} [normed_add_comm_group G'] [normed_space 𝕜 G'] + +variables {f f₀ f₁ g : E → F} +variables {f' f₀' f₁' g' : E →L[𝕜] F} +variables (e : E →L[𝕜] F) +variables {x : E} +variables {s t : set E} +variables {L L₁ L₂ : filter E} + +section clm_comp_apply +/-! ### Derivative of the pointwise composition/application of continuous linear maps -/ + +variables {H : Type*} [normed_add_comm_group H] [normed_space 𝕜 H] {c : E → G →L[𝕜] H} + {c' : E →L[𝕜] G →L[𝕜] H} {d : E → F →L[𝕜] G} {d' : E →L[𝕜] F →L[𝕜] G} {u : E → G} + {u' : E →L[𝕜] G} + +lemma has_strict_fderiv_at.clm_comp (hc : has_strict_fderiv_at c c' x) + (hd : has_strict_fderiv_at d d' x) : has_strict_fderiv_at (λ y, (c y).comp (d y)) + ((compL 𝕜 F G H (c x)).comp d' + ((compL 𝕜 F G H).flip (d x)).comp c') x := +(is_bounded_bilinear_map_comp.has_strict_fderiv_at (c x, d x)).comp x $ hc.prod hd + +lemma has_fderiv_within_at.clm_comp (hc : has_fderiv_within_at c c' s x) + (hd : has_fderiv_within_at d d' s x) : has_fderiv_within_at (λ y, (c y).comp (d y)) + ((compL 𝕜 F G H (c x)).comp d' + ((compL 𝕜 F G H).flip (d x)).comp c') s x := +(is_bounded_bilinear_map_comp.has_fderiv_at (c x, d x)).comp_has_fderiv_within_at x $ hc.prod hd + +lemma has_fderiv_at.clm_comp (hc : has_fderiv_at c c' x) + (hd : has_fderiv_at d d' x) : has_fderiv_at (λ y, (c y).comp (d y)) + ((compL 𝕜 F G H (c x)).comp d' + ((compL 𝕜 F G H).flip (d x)).comp c') x := +(is_bounded_bilinear_map_comp.has_fderiv_at (c x, d x)).comp x $ hc.prod hd + +lemma differentiable_within_at.clm_comp + (hc : differentiable_within_at 𝕜 c s x) (hd : differentiable_within_at 𝕜 d s x) : + differentiable_within_at 𝕜 (λ y, (c y).comp (d y)) s x := +(hc.has_fderiv_within_at.clm_comp hd.has_fderiv_within_at).differentiable_within_at + +lemma differentiable_at.clm_comp (hc : differentiable_at 𝕜 c x) + (hd : differentiable_at 𝕜 d x) : differentiable_at 𝕜 (λ y, (c y).comp (d y)) x := +(hc.has_fderiv_at.clm_comp hd.has_fderiv_at).differentiable_at + +lemma differentiable_on.clm_comp (hc : differentiable_on 𝕜 c s) (hd : differentiable_on 𝕜 d s) : + differentiable_on 𝕜 (λ y, (c y).comp (d y)) s := +λx hx, (hc x hx).clm_comp (hd x hx) + +lemma differentiable.clm_comp (hc : differentiable 𝕜 c) (hd : differentiable 𝕜 d) : + differentiable 𝕜 (λ y, (c y).comp (d y)) := +λx, (hc x).clm_comp (hd x) + +lemma fderiv_within_clm_comp (hxs : unique_diff_within_at 𝕜 s x) + (hc : differentiable_within_at 𝕜 c s x) (hd : differentiable_within_at 𝕜 d s x) : + fderiv_within 𝕜 (λ y, (c y).comp (d y)) s x = + (compL 𝕜 F G H (c x)).comp (fderiv_within 𝕜 d s x) + + ((compL 𝕜 F G H).flip (d x)).comp (fderiv_within 𝕜 c s x) := +(hc.has_fderiv_within_at.clm_comp hd.has_fderiv_within_at).fderiv_within hxs + +lemma fderiv_clm_comp (hc : differentiable_at 𝕜 c x) (hd : differentiable_at 𝕜 d x) : + fderiv 𝕜 (λ y, (c y).comp (d y)) x = + (compL 𝕜 F G H (c x)).comp (fderiv 𝕜 d x) + + ((compL 𝕜 F G H).flip (d x)).comp (fderiv 𝕜 c x) := +(hc.has_fderiv_at.clm_comp hd.has_fderiv_at).fderiv + +lemma has_strict_fderiv_at.clm_apply (hc : has_strict_fderiv_at c c' x) + (hu : has_strict_fderiv_at u u' x) : + has_strict_fderiv_at (λ y, (c y) (u y)) ((c x).comp u' + c'.flip (u x)) x := +(is_bounded_bilinear_map_apply.has_strict_fderiv_at (c x, u x)).comp x (hc.prod hu) + +lemma has_fderiv_within_at.clm_apply (hc : has_fderiv_within_at c c' s x) + (hu : has_fderiv_within_at u u' s x) : + has_fderiv_within_at (λ y, (c y) (u y)) ((c x).comp u' + c'.flip (u x)) s x := +(is_bounded_bilinear_map_apply.has_fderiv_at (c x, u x)).comp_has_fderiv_within_at x (hc.prod hu) + +lemma has_fderiv_at.clm_apply (hc : has_fderiv_at c c' x) (hu : has_fderiv_at u u' x) : + has_fderiv_at (λ y, (c y) (u y)) ((c x).comp u' + c'.flip (u x)) x := +(is_bounded_bilinear_map_apply.has_fderiv_at (c x, u x)).comp x (hc.prod hu) + +lemma differentiable_within_at.clm_apply + (hc : differentiable_within_at 𝕜 c s x) (hu : differentiable_within_at 𝕜 u s x) : + differentiable_within_at 𝕜 (λ y, (c y) (u y)) s x := +(hc.has_fderiv_within_at.clm_apply hu.has_fderiv_within_at).differentiable_within_at + +lemma differentiable_at.clm_apply (hc : differentiable_at 𝕜 c x) + (hu : differentiable_at 𝕜 u x) : differentiable_at 𝕜 (λ y, (c y) (u y)) x := +(hc.has_fderiv_at.clm_apply hu.has_fderiv_at).differentiable_at + +lemma differentiable_on.clm_apply (hc : differentiable_on 𝕜 c s) (hu : differentiable_on 𝕜 u s) : + differentiable_on 𝕜 (λ y, (c y) (u y)) s := +λx hx, (hc x hx).clm_apply (hu x hx) + +lemma differentiable.clm_apply (hc : differentiable 𝕜 c) (hu : differentiable 𝕜 u) : + differentiable 𝕜 (λ y, (c y) (u y)) := +λx, (hc x).clm_apply (hu x) + +lemma fderiv_within_clm_apply (hxs : unique_diff_within_at 𝕜 s x) + (hc : differentiable_within_at 𝕜 c s x) (hu : differentiable_within_at 𝕜 u s x) : + fderiv_within 𝕜 (λ y, (c y) (u y)) s x = + ((c x).comp (fderiv_within 𝕜 u s x) + (fderiv_within 𝕜 c s x).flip (u x)) := +(hc.has_fderiv_within_at.clm_apply hu.has_fderiv_within_at).fderiv_within hxs + +lemma fderiv_clm_apply (hc : differentiable_at 𝕜 c x) (hu : differentiable_at 𝕜 u x) : + fderiv 𝕜 (λ y, (c y) (u y)) x = ((c x).comp (fderiv 𝕜 u x) + (fderiv 𝕜 c x).flip (u x)) := +(hc.has_fderiv_at.clm_apply hu.has_fderiv_at).fderiv + +end clm_comp_apply + +section smul +/-! ### Derivative of the product of a scalar-valued function and a vector-valued function + +If `c` is a differentiable scalar-valued function and `f` is a differentiable vector-valued +function, then `λ x, c x • f x` is differentiable as well. Lemmas in this section works for +function `c` taking values in the base field, as well as in a normed algebra over the base +field: e.g., they work for `c : E → ℂ` and `f : E → F` provided that `F` is a complex +normed vector space. +-/ + +variables {𝕜' : Type*} [nontrivially_normed_field 𝕜'] [normed_algebra 𝕜 𝕜'] + [normed_space 𝕜' F] [is_scalar_tower 𝕜 𝕜' F] +variables {c : E → 𝕜'} {c' : E →L[𝕜] 𝕜'} + +theorem has_strict_fderiv_at.smul (hc : has_strict_fderiv_at c c' x) + (hf : has_strict_fderiv_at f f' x) : + has_strict_fderiv_at (λ y, c y • f y) (c x • f' + c'.smul_right (f x)) x := +(is_bounded_bilinear_map_smul.has_strict_fderiv_at (c x, f x)).comp x $ + hc.prod hf + +theorem has_fderiv_within_at.smul + (hc : has_fderiv_within_at c c' s x) (hf : has_fderiv_within_at f f' s x) : + has_fderiv_within_at (λ y, c y • f y) (c x • f' + c'.smul_right (f x)) s x := +(is_bounded_bilinear_map_smul.has_fderiv_at (c x, f x)).comp_has_fderiv_within_at x $ + hc.prod hf + +theorem has_fderiv_at.smul (hc : has_fderiv_at c c' x) (hf : has_fderiv_at f f' x) : + has_fderiv_at (λ y, c y • f y) (c x • f' + c'.smul_right (f x)) x := +(is_bounded_bilinear_map_smul.has_fderiv_at (c x, f x)).comp x $ + hc.prod hf + +lemma differentiable_within_at.smul + (hc : differentiable_within_at 𝕜 c s x) (hf : differentiable_within_at 𝕜 f s x) : + differentiable_within_at 𝕜 (λ y, c y • f y) s x := +(hc.has_fderiv_within_at.smul hf.has_fderiv_within_at).differentiable_within_at + +@[simp] lemma differentiable_at.smul (hc : differentiable_at 𝕜 c x) (hf : differentiable_at 𝕜 f x) : + differentiable_at 𝕜 (λ y, c y • f y) x := +(hc.has_fderiv_at.smul hf.has_fderiv_at).differentiable_at + +lemma differentiable_on.smul (hc : differentiable_on 𝕜 c s) (hf : differentiable_on 𝕜 f s) : + differentiable_on 𝕜 (λ y, c y • f y) s := +λx hx, (hc x hx).smul (hf x hx) + +@[simp] lemma differentiable.smul (hc : differentiable 𝕜 c) (hf : differentiable 𝕜 f) : + differentiable 𝕜 (λ y, c y • f y) := +λx, (hc x).smul (hf x) + +lemma fderiv_within_smul (hxs : unique_diff_within_at 𝕜 s x) + (hc : differentiable_within_at 𝕜 c s x) (hf : differentiable_within_at 𝕜 f s x) : + fderiv_within 𝕜 (λ y, c y • f y) s x = + c x • fderiv_within 𝕜 f s x + (fderiv_within 𝕜 c s x).smul_right (f x) := +(hc.has_fderiv_within_at.smul hf.has_fderiv_within_at).fderiv_within hxs + +lemma fderiv_smul (hc : differentiable_at 𝕜 c x) (hf : differentiable_at 𝕜 f x) : + fderiv 𝕜 (λ y, c y • f y) x = + c x • fderiv 𝕜 f x + (fderiv 𝕜 c x).smul_right (f x) := +(hc.has_fderiv_at.smul hf.has_fderiv_at).fderiv + +theorem has_strict_fderiv_at.smul_const (hc : has_strict_fderiv_at c c' x) (f : F) : + has_strict_fderiv_at (λ y, c y • f) (c'.smul_right f) x := +by simpa only [smul_zero, zero_add] using hc.smul (has_strict_fderiv_at_const f x) + +theorem has_fderiv_within_at.smul_const (hc : has_fderiv_within_at c c' s x) (f : F) : + has_fderiv_within_at (λ y, c y • f) (c'.smul_right f) s x := +by simpa only [smul_zero, zero_add] using hc.smul (has_fderiv_within_at_const f x s) + +theorem has_fderiv_at.smul_const (hc : has_fderiv_at c c' x) (f : F) : + has_fderiv_at (λ y, c y • f) (c'.smul_right f) x := +by simpa only [smul_zero, zero_add] using hc.smul (has_fderiv_at_const f x) + +lemma differentiable_within_at.smul_const + (hc : differentiable_within_at 𝕜 c s x) (f : F) : + differentiable_within_at 𝕜 (λ y, c y • f) s x := +(hc.has_fderiv_within_at.smul_const f).differentiable_within_at + +lemma differentiable_at.smul_const (hc : differentiable_at 𝕜 c x) (f : F) : + differentiable_at 𝕜 (λ y, c y • f) x := +(hc.has_fderiv_at.smul_const f).differentiable_at + +lemma differentiable_on.smul_const (hc : differentiable_on 𝕜 c s) (f : F) : + differentiable_on 𝕜 (λ y, c y • f) s := +λx hx, (hc x hx).smul_const f + +lemma differentiable.smul_const (hc : differentiable 𝕜 c) (f : F) : + differentiable 𝕜 (λ y, c y • f) := +λx, (hc x).smul_const f + +lemma fderiv_within_smul_const (hxs : unique_diff_within_at 𝕜 s x) + (hc : differentiable_within_at 𝕜 c s x) (f : F) : + fderiv_within 𝕜 (λ y, c y • f) s x = + (fderiv_within 𝕜 c s x).smul_right f := +(hc.has_fderiv_within_at.smul_const f).fderiv_within hxs + +lemma fderiv_smul_const (hc : differentiable_at 𝕜 c x) (f : F) : + fderiv 𝕜 (λ y, c y • f) x = (fderiv 𝕜 c x).smul_right f := +(hc.has_fderiv_at.smul_const f).fderiv + +end smul + +section mul +/-! ### Derivative of the product of two functions -/ + +variables {𝔸 𝔸' : Type*} [normed_ring 𝔸] [normed_comm_ring 𝔸'] [normed_algebra 𝕜 𝔸] + [normed_algebra 𝕜 𝔸'] {a b : E → 𝔸} {a' b' : E →L[𝕜] 𝔸} {c d : E → 𝔸'} {c' d' : E →L[𝕜] 𝔸'} + +theorem has_strict_fderiv_at.mul' {x : E} (ha : has_strict_fderiv_at a a' x) + (hb : has_strict_fderiv_at b b' x) : + has_strict_fderiv_at (λ y, a y * b y) (a x • b' + a'.smul_right (b x)) x := +((continuous_linear_map.mul 𝕜 𝔸).is_bounded_bilinear_map.has_strict_fderiv_at (a x, b x)).comp x + (ha.prod hb) + +theorem has_strict_fderiv_at.mul + (hc : has_strict_fderiv_at c c' x) (hd : has_strict_fderiv_at d d' x) : + has_strict_fderiv_at (λ y, c y * d y) (c x • d' + d x • c') x := +by { convert hc.mul' hd, ext z, apply mul_comm } + +theorem has_fderiv_within_at.mul' + (ha : has_fderiv_within_at a a' s x) (hb : has_fderiv_within_at b b' s x) : + has_fderiv_within_at (λ y, a y * b y) (a x • b' + a'.smul_right (b x)) s x := +((continuous_linear_map.mul 𝕜 𝔸).is_bounded_bilinear_map.has_fderiv_at + (a x, b x)).comp_has_fderiv_within_at x (ha.prod hb) + +theorem has_fderiv_within_at.mul + (hc : has_fderiv_within_at c c' s x) (hd : has_fderiv_within_at d d' s x) : + has_fderiv_within_at (λ y, c y * d y) (c x • d' + d x • c') s x := +by { convert hc.mul' hd, ext z, apply mul_comm } + +theorem has_fderiv_at.mul' + (ha : has_fderiv_at a a' x) (hb : has_fderiv_at b b' x) : + has_fderiv_at (λ y, a y * b y) (a x • b' + a'.smul_right (b x)) x := +((continuous_linear_map.mul 𝕜 𝔸).is_bounded_bilinear_map.has_fderiv_at (a x, b x)).comp x + (ha.prod hb) + +theorem has_fderiv_at.mul (hc : has_fderiv_at c c' x) (hd : has_fderiv_at d d' x) : + has_fderiv_at (λ y, c y * d y) (c x • d' + d x • c') x := +by { convert hc.mul' hd, ext z, apply mul_comm } + +lemma differentiable_within_at.mul + (ha : differentiable_within_at 𝕜 a s x) (hb : differentiable_within_at 𝕜 b s x) : + differentiable_within_at 𝕜 (λ y, a y * b y) s x := +(ha.has_fderiv_within_at.mul' hb.has_fderiv_within_at).differentiable_within_at + +@[simp] lemma differentiable_at.mul (ha : differentiable_at 𝕜 a x) (hb : differentiable_at 𝕜 b x) : + differentiable_at 𝕜 (λ y, a y * b y) x := +(ha.has_fderiv_at.mul' hb.has_fderiv_at).differentiable_at + +lemma differentiable_on.mul (ha : differentiable_on 𝕜 a s) (hb : differentiable_on 𝕜 b s) : + differentiable_on 𝕜 (λ y, a y * b y) s := +λx hx, (ha x hx).mul (hb x hx) + +@[simp] lemma differentiable.mul (ha : differentiable 𝕜 a) (hb : differentiable 𝕜 b) : + differentiable 𝕜 (λ y, a y * b y) := +λx, (ha x).mul (hb x) + +lemma differentiable_within_at.pow (ha : differentiable_within_at 𝕜 a s x) : + ∀ n : ℕ, differentiable_within_at 𝕜 (λ x, a x ^ n) s x +| 0 := by simp only [pow_zero, differentiable_within_at_const] +| (n + 1) := by simp only [pow_succ, differentiable_within_at.pow n, ha.mul] + +@[simp] lemma differentiable_at.pow (ha : differentiable_at 𝕜 a x) (n : ℕ) : + differentiable_at 𝕜 (λ x, a x ^ n) x := +differentiable_within_at_univ.mp $ ha.differentiable_within_at.pow n + +lemma differentiable_on.pow (ha : differentiable_on 𝕜 a s) (n : ℕ) : + differentiable_on 𝕜 (λ x, a x ^ n) s := +λ x h, (ha x h).pow n + +@[simp] lemma differentiable.pow (ha : differentiable 𝕜 a) (n : ℕ) : + differentiable 𝕜 (λ x, a x ^ n) := +λx, (ha x).pow n + +lemma fderiv_within_mul' (hxs : unique_diff_within_at 𝕜 s x) + (ha : differentiable_within_at 𝕜 a s x) (hb : differentiable_within_at 𝕜 b s x) : + fderiv_within 𝕜 (λ y, a y * b y) s x = + a x • fderiv_within 𝕜 b s x + (fderiv_within 𝕜 a s x).smul_right (b x) := +(ha.has_fderiv_within_at.mul' hb.has_fderiv_within_at).fderiv_within hxs + +lemma fderiv_within_mul (hxs : unique_diff_within_at 𝕜 s x) + (hc : differentiable_within_at 𝕜 c s x) (hd : differentiable_within_at 𝕜 d s x) : + fderiv_within 𝕜 (λ y, c y * d y) s x = + c x • fderiv_within 𝕜 d s x + d x • fderiv_within 𝕜 c s x := +(hc.has_fderiv_within_at.mul hd.has_fderiv_within_at).fderiv_within hxs + +lemma fderiv_mul' (ha : differentiable_at 𝕜 a x) (hb : differentiable_at 𝕜 b x) : + fderiv 𝕜 (λ y, a y * b y) x = + a x • fderiv 𝕜 b x + (fderiv 𝕜 a x).smul_right (b x) := +(ha.has_fderiv_at.mul' hb.has_fderiv_at).fderiv + +lemma fderiv_mul (hc : differentiable_at 𝕜 c x) (hd : differentiable_at 𝕜 d x) : + fderiv 𝕜 (λ y, c y * d y) x = + c x • fderiv 𝕜 d x + d x • fderiv 𝕜 c x := +(hc.has_fderiv_at.mul hd.has_fderiv_at).fderiv + +theorem has_strict_fderiv_at.mul_const' (ha : has_strict_fderiv_at a a' x) (b : 𝔸) : + has_strict_fderiv_at (λ y, a y * b) (a'.smul_right b) x := +(((continuous_linear_map.mul 𝕜 𝔸).flip b).has_strict_fderiv_at).comp x ha + +theorem has_strict_fderiv_at.mul_const (hc : has_strict_fderiv_at c c' x) (d : 𝔸') : + has_strict_fderiv_at (λ y, c y * d) (d • c') x := +by { convert hc.mul_const' d, ext z, apply mul_comm } + +theorem has_fderiv_within_at.mul_const' (ha : has_fderiv_within_at a a' s x) (b : 𝔸) : + has_fderiv_within_at (λ y, a y * b) (a'.smul_right b) s x := +(((continuous_linear_map.mul 𝕜 𝔸).flip b).has_fderiv_at).comp_has_fderiv_within_at x ha + +theorem has_fderiv_within_at.mul_const (hc : has_fderiv_within_at c c' s x) (d : 𝔸') : + has_fderiv_within_at (λ y, c y * d) (d • c') s x := +by { convert hc.mul_const' d, ext z, apply mul_comm } + +theorem has_fderiv_at.mul_const' (ha : has_fderiv_at a a' x) (b : 𝔸) : + has_fderiv_at (λ y, a y * b) (a'.smul_right b) x := +(((continuous_linear_map.mul 𝕜 𝔸).flip b).has_fderiv_at).comp x ha + +theorem has_fderiv_at.mul_const (hc : has_fderiv_at c c' x) (d : 𝔸') : + has_fderiv_at (λ y, c y * d) (d • c') x := +by { convert hc.mul_const' d, ext z, apply mul_comm } + +lemma differentiable_within_at.mul_const + (ha : differentiable_within_at 𝕜 a s x) (b : 𝔸) : + differentiable_within_at 𝕜 (λ y, a y * b) s x := +(ha.has_fderiv_within_at.mul_const' b).differentiable_within_at + +lemma differentiable_at.mul_const (ha : differentiable_at 𝕜 a x) (b : 𝔸) : + differentiable_at 𝕜 (λ y, a y * b) x := +(ha.has_fderiv_at.mul_const' b).differentiable_at + +lemma differentiable_on.mul_const (ha : differentiable_on 𝕜 a s) (b : 𝔸) : + differentiable_on 𝕜 (λ y, a y * b) s := +λx hx, (ha x hx).mul_const b + +lemma differentiable.mul_const (ha : differentiable 𝕜 a) (b : 𝔸) : + differentiable 𝕜 (λ y, a y * b) := +λx, (ha x).mul_const b + +lemma fderiv_within_mul_const' (hxs : unique_diff_within_at 𝕜 s x) + (ha : differentiable_within_at 𝕜 a s x) (b : 𝔸) : + fderiv_within 𝕜 (λ y, a y * b) s x = (fderiv_within 𝕜 a s x).smul_right b := +(ha.has_fderiv_within_at.mul_const' b).fderiv_within hxs + +lemma fderiv_within_mul_const (hxs : unique_diff_within_at 𝕜 s x) + (hc : differentiable_within_at 𝕜 c s x) (d : 𝔸') : + fderiv_within 𝕜 (λ y, c y * d) s x = d • fderiv_within 𝕜 c s x := +(hc.has_fderiv_within_at.mul_const d).fderiv_within hxs + +lemma fderiv_mul_const' (ha : differentiable_at 𝕜 a x) (b : 𝔸) : + fderiv 𝕜 (λ y, a y * b) x = (fderiv 𝕜 a x).smul_right b := +(ha.has_fderiv_at.mul_const' b).fderiv + +lemma fderiv_mul_const (hc : differentiable_at 𝕜 c x) (d : 𝔸') : + fderiv 𝕜 (λ y, c y * d) x = d • fderiv 𝕜 c x := +(hc.has_fderiv_at.mul_const d).fderiv + +theorem has_strict_fderiv_at.const_mul (ha : has_strict_fderiv_at a a' x) (b : 𝔸) : + has_strict_fderiv_at (λ y, b * a y) (b • a') x := +(((continuous_linear_map.mul 𝕜 𝔸) b).has_strict_fderiv_at).comp x ha + +theorem has_fderiv_within_at.const_mul + (ha : has_fderiv_within_at a a' s x) (b : 𝔸) : + has_fderiv_within_at (λ y, b * a y) (b • a') s x := +(((continuous_linear_map.mul 𝕜 𝔸) b).has_fderiv_at).comp_has_fderiv_within_at x ha + +theorem has_fderiv_at.const_mul (ha : has_fderiv_at a a' x) (b : 𝔸) : + has_fderiv_at (λ y, b * a y) (b • a') x := +(((continuous_linear_map.mul 𝕜 𝔸) b).has_fderiv_at).comp x ha + +lemma differentiable_within_at.const_mul + (ha : differentiable_within_at 𝕜 a s x) (b : 𝔸) : + differentiable_within_at 𝕜 (λ y, b * a y) s x := +(ha.has_fderiv_within_at.const_mul b).differentiable_within_at + +lemma differentiable_at.const_mul (ha : differentiable_at 𝕜 a x) (b : 𝔸) : + differentiable_at 𝕜 (λ y, b * a y) x := +(ha.has_fderiv_at.const_mul b).differentiable_at + +lemma differentiable_on.const_mul (ha : differentiable_on 𝕜 a s) (b : 𝔸) : + differentiable_on 𝕜 (λ y, b * a y) s := +λx hx, (ha x hx).const_mul b + +lemma differentiable.const_mul (ha : differentiable 𝕜 a) (b : 𝔸) : + differentiable 𝕜 (λ y, b * a y) := +λx, (ha x).const_mul b + +lemma fderiv_within_const_mul (hxs : unique_diff_within_at 𝕜 s x) + (ha : differentiable_within_at 𝕜 a s x) (b : 𝔸) : + fderiv_within 𝕜 (λ y, b * a y) s x = b • fderiv_within 𝕜 a s x := +(ha.has_fderiv_within_at.const_mul b).fderiv_within hxs + +lemma fderiv_const_mul (ha : differentiable_at 𝕜 a x) (b : 𝔸) : + fderiv 𝕜 (λ y, b * a y) x = b • fderiv 𝕜 a x := +(ha.has_fderiv_at.const_mul b).fderiv + +end mul + +section algebra_inverse +variables {R : Type*} [normed_ring R] [normed_algebra 𝕜 R] [complete_space R] +open normed_ring continuous_linear_map ring + +/-- At an invertible element `x` of a normed algebra `R`, the Fréchet derivative of the inversion +operation is the linear map `λ t, - x⁻¹ * t * x⁻¹`. -/ +lemma has_fderiv_at_ring_inverse (x : Rˣ) : + has_fderiv_at ring.inverse (-mul_left_right 𝕜 R ↑x⁻¹ ↑x⁻¹) x := +begin + have h_is_o : (λ (t : R), inverse (↑x + t) - ↑x⁻¹ + ↑x⁻¹ * t * ↑x⁻¹) =o[𝓝 0] (λ (t : R), t), + { refine (inverse_add_norm_diff_second_order x).trans_is_o ((is_o_norm_norm).mp _), + simp only [norm_pow, norm_norm], + have h12 : 1 < 2 := by norm_num, + convert (asymptotics.is_o_pow_pow h12).comp_tendsto tendsto_norm_zero, + ext, simp }, + have h_lim : tendsto (λ (y:R), y - x) (𝓝 x) (𝓝 0), + { refine tendsto_zero_iff_norm_tendsto_zero.mpr _, + exact tendsto_iff_norm_tendsto_zero.mp tendsto_id }, + simp only [has_fderiv_at, has_fderiv_at_filter], + convert h_is_o.comp_tendsto h_lim, + ext y, + simp only [coe_comp', function.comp_app, mul_left_right_apply, neg_apply, inverse_unit x, + units.inv_mul, add_sub_cancel'_right, mul_sub, sub_mul, one_mul, sub_neg_eq_add] +end + +lemma differentiable_at_inverse (x : Rˣ) : differentiable_at 𝕜 (@ring.inverse R _) x := +(has_fderiv_at_ring_inverse x).differentiable_at + +lemma fderiv_inverse (x : Rˣ) : + fderiv 𝕜 (@ring.inverse R _) x = - mul_left_right 𝕜 R ↑x⁻¹ ↑x⁻¹ := +(has_fderiv_at_ring_inverse x).fderiv + +end algebra_inverse + +end diff --git a/src/analysis/calculus/fderiv/prod.lean b/src/analysis/calculus/fderiv/prod.lean new file mode 100644 index 0000000000000..01b3a4f00d183 --- /dev/null +++ b/src/analysis/calculus/fderiv/prod.lean @@ -0,0 +1,383 @@ +/- +Copyright (c) 2019 Jeremy Avigad. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jeremy Avigad, Sébastien Gouëzel, Yury Kudryashov +-/ +import analysis.calculus.fderiv.linear +import analysis.calculus.fderiv.comp + +/-! +# Derivative of the cartesian product of functions + +For detailed documentation of the Fréchet derivative, +see the module docstring of `analysis/calculus/fderiv/basic.lean`. + +This file contains the usual formulas (and existence assertions) for the derivative of +cartesian products of functions, and functions into Pi-types. +-/ + +open filter asymptotics continuous_linear_map set metric +open_locale topology classical nnreal filter asymptotics ennreal + +noncomputable theory + + +section + +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] +variables {E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] +variables {F : Type*} [normed_add_comm_group F] [normed_space 𝕜 F] +variables {G : Type*} [normed_add_comm_group G] [normed_space 𝕜 G] +variables {G' : Type*} [normed_add_comm_group G'] [normed_space 𝕜 G'] + +variables {f f₀ f₁ g : E → F} +variables {f' f₀' f₁' g' : E →L[𝕜] F} +variables (e : E →L[𝕜] F) +variables {x : E} +variables {s t : set E} +variables {L L₁ L₂ : filter E} + +section cartesian_product +/-! ### Derivative of the cartesian product of two functions -/ + +section prod +variables {f₂ : E → G} {f₂' : E →L[𝕜] G} + +protected lemma has_strict_fderiv_at.prod + (hf₁ : has_strict_fderiv_at f₁ f₁' x) (hf₂ : has_strict_fderiv_at f₂ f₂' x) : + has_strict_fderiv_at (λx, (f₁ x, f₂ x)) (f₁'.prod f₂') x := +hf₁.prod_left hf₂ + +lemma has_fderiv_at_filter.prod + (hf₁ : has_fderiv_at_filter f₁ f₁' x L) (hf₂ : has_fderiv_at_filter f₂ f₂' x L) : + has_fderiv_at_filter (λx, (f₁ x, f₂ x)) (f₁'.prod f₂') x L := +hf₁.prod_left hf₂ + +lemma has_fderiv_within_at.prod + (hf₁ : has_fderiv_within_at f₁ f₁' s x) (hf₂ : has_fderiv_within_at f₂ f₂' s x) : + has_fderiv_within_at (λx, (f₁ x, f₂ x)) (f₁'.prod f₂') s x := +hf₁.prod hf₂ + +lemma has_fderiv_at.prod (hf₁ : has_fderiv_at f₁ f₁' x) (hf₂ : has_fderiv_at f₂ f₂' x) : + has_fderiv_at (λx, (f₁ x, f₂ x)) (f₁'.prod f₂') x := +hf₁.prod hf₂ + +lemma has_fderiv_at_prod_mk_left (e₀ : E) (f₀ : F) : + has_fderiv_at (λ e : E, (e, f₀)) (inl 𝕜 E F) e₀ := +(has_fderiv_at_id e₀).prod (has_fderiv_at_const f₀ e₀) + +lemma has_fderiv_at_prod_mk_right (e₀ : E) (f₀ : F) : + has_fderiv_at (λ f : F, (e₀, f)) (inr 𝕜 E F) f₀ := +(has_fderiv_at_const e₀ f₀).prod (has_fderiv_at_id f₀) + +lemma differentiable_within_at.prod + (hf₁ : differentiable_within_at 𝕜 f₁ s x) (hf₂ : differentiable_within_at 𝕜 f₂ s x) : + differentiable_within_at 𝕜 (λx:E, (f₁ x, f₂ x)) s x := +(hf₁.has_fderiv_within_at.prod hf₂.has_fderiv_within_at).differentiable_within_at + +@[simp] +lemma differentiable_at.prod (hf₁ : differentiable_at 𝕜 f₁ x) (hf₂ : differentiable_at 𝕜 f₂ x) : + differentiable_at 𝕜 (λx:E, (f₁ x, f₂ x)) x := +(hf₁.has_fderiv_at.prod hf₂.has_fderiv_at).differentiable_at + +lemma differentiable_on.prod (hf₁ : differentiable_on 𝕜 f₁ s) (hf₂ : differentiable_on 𝕜 f₂ s) : + differentiable_on 𝕜 (λx:E, (f₁ x, f₂ x)) s := +λx hx, differentiable_within_at.prod (hf₁ x hx) (hf₂ x hx) + +@[simp] +lemma differentiable.prod (hf₁ : differentiable 𝕜 f₁) (hf₂ : differentiable 𝕜 f₂) : + differentiable 𝕜 (λx:E, (f₁ x, f₂ x)) := +λ x, differentiable_at.prod (hf₁ x) (hf₂ x) + +lemma differentiable_at.fderiv_prod + (hf₁ : differentiable_at 𝕜 f₁ x) (hf₂ : differentiable_at 𝕜 f₂ x) : + fderiv 𝕜 (λx:E, (f₁ x, f₂ x)) x = (fderiv 𝕜 f₁ x).prod (fderiv 𝕜 f₂ x) := +(hf₁.has_fderiv_at.prod hf₂.has_fderiv_at).fderiv + +lemma differentiable_at.fderiv_within_prod + (hf₁ : differentiable_within_at 𝕜 f₁ s x) (hf₂ : differentiable_within_at 𝕜 f₂ s x) + (hxs : unique_diff_within_at 𝕜 s x) : + fderiv_within 𝕜 (λx:E, (f₁ x, f₂ x)) s x = + (fderiv_within 𝕜 f₁ s x).prod (fderiv_within 𝕜 f₂ s x) := +(hf₁.has_fderiv_within_at.prod hf₂.has_fderiv_within_at).fderiv_within hxs + +end prod + +section fst + +variables {f₂ : E → F × G} {f₂' : E →L[𝕜] F × G} {p : E × F} + +lemma has_strict_fderiv_at_fst : has_strict_fderiv_at (@prod.fst E F) (fst 𝕜 E F) p := +(fst 𝕜 E F).has_strict_fderiv_at + +protected lemma has_strict_fderiv_at.fst (h : has_strict_fderiv_at f₂ f₂' x) : + has_strict_fderiv_at (λ x, (f₂ x).1) ((fst 𝕜 F G).comp f₂') x := +has_strict_fderiv_at_fst.comp x h + +lemma has_fderiv_at_filter_fst {L : filter (E × F)} : + has_fderiv_at_filter (@prod.fst E F) (fst 𝕜 E F) p L := +(fst 𝕜 E F).has_fderiv_at_filter + +protected lemma has_fderiv_at_filter.fst (h : has_fderiv_at_filter f₂ f₂' x L) : + has_fderiv_at_filter (λ x, (f₂ x).1) ((fst 𝕜 F G).comp f₂') x L := +has_fderiv_at_filter_fst.comp x h tendsto_map + +lemma has_fderiv_at_fst : has_fderiv_at (@prod.fst E F) (fst 𝕜 E F) p := +has_fderiv_at_filter_fst + +protected lemma has_fderiv_at.fst (h : has_fderiv_at f₂ f₂' x) : + has_fderiv_at (λ x, (f₂ x).1) ((fst 𝕜 F G).comp f₂') x := +h.fst + +lemma has_fderiv_within_at_fst {s : set (E × F)} : + has_fderiv_within_at (@prod.fst E F) (fst 𝕜 E F) s p := +has_fderiv_at_filter_fst + +protected lemma has_fderiv_within_at.fst (h : has_fderiv_within_at f₂ f₂' s x) : + has_fderiv_within_at (λ x, (f₂ x).1) ((fst 𝕜 F G).comp f₂') s x := +h.fst + +lemma differentiable_at_fst : differentiable_at 𝕜 prod.fst p := +has_fderiv_at_fst.differentiable_at + +@[simp] protected lemma differentiable_at.fst (h : differentiable_at 𝕜 f₂ x) : + differentiable_at 𝕜 (λ x, (f₂ x).1) x := +differentiable_at_fst.comp x h + +lemma differentiable_fst : differentiable 𝕜 (prod.fst : E × F → E) := +λ x, differentiable_at_fst + +@[simp] protected lemma differentiable.fst (h : differentiable 𝕜 f₂) : + differentiable 𝕜 (λ x, (f₂ x).1) := +differentiable_fst.comp h + +lemma differentiable_within_at_fst {s : set (E × F)} : differentiable_within_at 𝕜 prod.fst s p := +differentiable_at_fst.differentiable_within_at + +protected lemma differentiable_within_at.fst (h : differentiable_within_at 𝕜 f₂ s x) : + differentiable_within_at 𝕜 (λ x, (f₂ x).1) s x := +differentiable_at_fst.comp_differentiable_within_at x h + +lemma differentiable_on_fst {s : set (E × F)} : differentiable_on 𝕜 prod.fst s := +differentiable_fst.differentiable_on + +protected lemma differentiable_on.fst (h : differentiable_on 𝕜 f₂ s) : + differentiable_on 𝕜 (λ x, (f₂ x).1) s := +differentiable_fst.comp_differentiable_on h + +lemma fderiv_fst : fderiv 𝕜 prod.fst p = fst 𝕜 E F := has_fderiv_at_fst.fderiv + +lemma fderiv.fst (h : differentiable_at 𝕜 f₂ x) : + fderiv 𝕜 (λ x, (f₂ x).1) x = (fst 𝕜 F G).comp (fderiv 𝕜 f₂ x) := +h.has_fderiv_at.fst.fderiv + +lemma fderiv_within_fst {s : set (E × F)} (hs : unique_diff_within_at 𝕜 s p) : + fderiv_within 𝕜 prod.fst s p = fst 𝕜 E F := +has_fderiv_within_at_fst.fderiv_within hs + +lemma fderiv_within.fst (hs : unique_diff_within_at 𝕜 s x) (h : differentiable_within_at 𝕜 f₂ s x) : + fderiv_within 𝕜 (λ x, (f₂ x).1) s x = (fst 𝕜 F G).comp (fderiv_within 𝕜 f₂ s x) := +h.has_fderiv_within_at.fst.fderiv_within hs + +end fst + +section snd + +variables {f₂ : E → F × G} {f₂' : E →L[𝕜] F × G} {p : E × F} + +lemma has_strict_fderiv_at_snd : has_strict_fderiv_at (@prod.snd E F) (snd 𝕜 E F) p := +(snd 𝕜 E F).has_strict_fderiv_at + +protected lemma has_strict_fderiv_at.snd (h : has_strict_fderiv_at f₂ f₂' x) : + has_strict_fderiv_at (λ x, (f₂ x).2) ((snd 𝕜 F G).comp f₂') x := +has_strict_fderiv_at_snd.comp x h + +lemma has_fderiv_at_filter_snd {L : filter (E × F)} : + has_fderiv_at_filter (@prod.snd E F) (snd 𝕜 E F) p L := +(snd 𝕜 E F).has_fderiv_at_filter + +protected lemma has_fderiv_at_filter.snd (h : has_fderiv_at_filter f₂ f₂' x L) : + has_fderiv_at_filter (λ x, (f₂ x).2) ((snd 𝕜 F G).comp f₂') x L := +has_fderiv_at_filter_snd.comp x h tendsto_map + +lemma has_fderiv_at_snd : has_fderiv_at (@prod.snd E F) (snd 𝕜 E F) p := +has_fderiv_at_filter_snd + +protected lemma has_fderiv_at.snd (h : has_fderiv_at f₂ f₂' x) : + has_fderiv_at (λ x, (f₂ x).2) ((snd 𝕜 F G).comp f₂') x := +h.snd + +lemma has_fderiv_within_at_snd {s : set (E × F)} : + has_fderiv_within_at (@prod.snd E F) (snd 𝕜 E F) s p := +has_fderiv_at_filter_snd + +protected lemma has_fderiv_within_at.snd (h : has_fderiv_within_at f₂ f₂' s x) : + has_fderiv_within_at (λ x, (f₂ x).2) ((snd 𝕜 F G).comp f₂') s x := +h.snd + +lemma differentiable_at_snd : differentiable_at 𝕜 prod.snd p := +has_fderiv_at_snd.differentiable_at + +@[simp] protected lemma differentiable_at.snd (h : differentiable_at 𝕜 f₂ x) : + differentiable_at 𝕜 (λ x, (f₂ x).2) x := +differentiable_at_snd.comp x h + +lemma differentiable_snd : differentiable 𝕜 (prod.snd : E × F → F) := +λ x, differentiable_at_snd + +@[simp] protected lemma differentiable.snd (h : differentiable 𝕜 f₂) : + differentiable 𝕜 (λ x, (f₂ x).2) := +differentiable_snd.comp h + +lemma differentiable_within_at_snd {s : set (E × F)} : differentiable_within_at 𝕜 prod.snd s p := +differentiable_at_snd.differentiable_within_at + +protected lemma differentiable_within_at.snd (h : differentiable_within_at 𝕜 f₂ s x) : + differentiable_within_at 𝕜 (λ x, (f₂ x).2) s x := +differentiable_at_snd.comp_differentiable_within_at x h + +lemma differentiable_on_snd {s : set (E × F)} : differentiable_on 𝕜 prod.snd s := +differentiable_snd.differentiable_on + +protected lemma differentiable_on.snd (h : differentiable_on 𝕜 f₂ s) : + differentiable_on 𝕜 (λ x, (f₂ x).2) s := +differentiable_snd.comp_differentiable_on h + +lemma fderiv_snd : fderiv 𝕜 prod.snd p = snd 𝕜 E F := has_fderiv_at_snd.fderiv + +lemma fderiv.snd (h : differentiable_at 𝕜 f₂ x) : + fderiv 𝕜 (λ x, (f₂ x).2) x = (snd 𝕜 F G).comp (fderiv 𝕜 f₂ x) := +h.has_fderiv_at.snd.fderiv + +lemma fderiv_within_snd {s : set (E × F)} (hs : unique_diff_within_at 𝕜 s p) : + fderiv_within 𝕜 prod.snd s p = snd 𝕜 E F := +has_fderiv_within_at_snd.fderiv_within hs + +lemma fderiv_within.snd (hs : unique_diff_within_at 𝕜 s x) (h : differentiable_within_at 𝕜 f₂ s x) : + fderiv_within 𝕜 (λ x, (f₂ x).2) s x = (snd 𝕜 F G).comp (fderiv_within 𝕜 f₂ s x) := +h.has_fderiv_within_at.snd.fderiv_within hs + +end snd + +section prod_map + +variables {f₂ : G → G'} {f₂' : G →L[𝕜] G'} {y : G} (p : E × G) + +protected theorem has_strict_fderiv_at.prod_map (hf : has_strict_fderiv_at f f' p.1) + (hf₂ : has_strict_fderiv_at f₂ f₂' p.2) : + has_strict_fderiv_at (prod.map f f₂) (f'.prod_map f₂') p := +(hf.comp p has_strict_fderiv_at_fst).prod (hf₂.comp p has_strict_fderiv_at_snd) + +protected theorem has_fderiv_at.prod_map (hf : has_fderiv_at f f' p.1) + (hf₂ : has_fderiv_at f₂ f₂' p.2) : + has_fderiv_at (prod.map f f₂) (f'.prod_map f₂') p := +(hf.comp p has_fderiv_at_fst).prod (hf₂.comp p has_fderiv_at_snd) + +@[simp] protected theorem differentiable_at.prod_map (hf : differentiable_at 𝕜 f p.1) + (hf₂ : differentiable_at 𝕜 f₂ p.2) : + differentiable_at 𝕜 (λ p : E × G, (f p.1, f₂ p.2)) p := +(hf.comp p differentiable_at_fst).prod (hf₂.comp p differentiable_at_snd) + +end prod_map + +section pi + +/-! +### Derivatives of functions `f : E → Π i, F' i` + +In this section we formulate `has_*fderiv*_pi` theorems as `iff`s, and provide two versions of each +theorem: + +* the version without `'` deals with `φ : Π i, E → F' i` and `φ' : Π i, E →L[𝕜] F' i` + and is designed to deduce differentiability of `λ x i, φ i x` from differentiability + of each `φ i`; +* the version with `'` deals with `Φ : E → Π i, F' i` and `Φ' : E →L[𝕜] Π i, F' i` + and is designed to deduce differentiability of the components `λ x, Φ x i` from + differentiability of `Φ`. +-/ + +variables {ι : Type*} [fintype ι] {F' : ι → Type*} [Π i, normed_add_comm_group (F' i)] + [Π i, normed_space 𝕜 (F' i)] {φ : Π i, E → F' i} {φ' : Π i, E →L[𝕜] F' i} + {Φ : E → Π i, F' i} {Φ' : E →L[𝕜] Π i, F' i} + +@[simp] lemma has_strict_fderiv_at_pi' : + has_strict_fderiv_at Φ Φ' x ↔ + ∀ i, has_strict_fderiv_at (λ x, Φ x i) ((proj i).comp Φ') x := +begin + simp only [has_strict_fderiv_at, continuous_linear_map.coe_pi], + exact is_o_pi +end + +@[simp] lemma has_strict_fderiv_at_pi : + has_strict_fderiv_at (λ x i, φ i x) (continuous_linear_map.pi φ') x ↔ + ∀ i, has_strict_fderiv_at (φ i) (φ' i) x := +has_strict_fderiv_at_pi' + +@[simp] lemma has_fderiv_at_filter_pi' : + has_fderiv_at_filter Φ Φ' x L ↔ + ∀ i, has_fderiv_at_filter (λ x, Φ x i) ((proj i).comp Φ') x L := +begin + simp only [has_fderiv_at_filter, continuous_linear_map.coe_pi], + exact is_o_pi +end + +lemma has_fderiv_at_filter_pi : + has_fderiv_at_filter (λ x i, φ i x) (continuous_linear_map.pi φ') x L ↔ + ∀ i, has_fderiv_at_filter (φ i) (φ' i) x L := +has_fderiv_at_filter_pi' + +@[simp] lemma has_fderiv_at_pi' : + has_fderiv_at Φ Φ' x ↔ + ∀ i, has_fderiv_at (λ x, Φ x i) ((proj i).comp Φ') x := +has_fderiv_at_filter_pi' + +lemma has_fderiv_at_pi : + has_fderiv_at (λ x i, φ i x) (continuous_linear_map.pi φ') x ↔ + ∀ i, has_fderiv_at (φ i) (φ' i) x := +has_fderiv_at_filter_pi + +@[simp] lemma has_fderiv_within_at_pi' : + has_fderiv_within_at Φ Φ' s x ↔ + ∀ i, has_fderiv_within_at (λ x, Φ x i) ((proj i).comp Φ') s x := +has_fderiv_at_filter_pi' + +lemma has_fderiv_within_at_pi : + has_fderiv_within_at (λ x i, φ i x) (continuous_linear_map.pi φ') s x ↔ + ∀ i, has_fderiv_within_at (φ i) (φ' i) s x := +has_fderiv_at_filter_pi + +@[simp] lemma differentiable_within_at_pi : + differentiable_within_at 𝕜 Φ s x ↔ + ∀ i, differentiable_within_at 𝕜 (λ x, Φ x i) s x := +⟨λ h i, (has_fderiv_within_at_pi'.1 h.has_fderiv_within_at i).differentiable_within_at, + λ h, (has_fderiv_within_at_pi.2 (λ i, (h i).has_fderiv_within_at)).differentiable_within_at⟩ + +@[simp] lemma differentiable_at_pi : + differentiable_at 𝕜 Φ x ↔ ∀ i, differentiable_at 𝕜 (λ x, Φ x i) x := +⟨λ h i, (has_fderiv_at_pi'.1 h.has_fderiv_at i).differentiable_at, + λ h, (has_fderiv_at_pi.2 (λ i, (h i).has_fderiv_at)).differentiable_at⟩ + +lemma differentiable_on_pi : + differentiable_on 𝕜 Φ s ↔ ∀ i, differentiable_on 𝕜 (λ x, Φ x i) s := +⟨λ h i x hx, differentiable_within_at_pi.1 (h x hx) i, + λ h x hx, differentiable_within_at_pi.2 (λ i, h i x hx)⟩ + +lemma differentiable_pi : + differentiable 𝕜 Φ ↔ ∀ i, differentiable 𝕜 (λ x, Φ x i) := +⟨λ h i x, differentiable_at_pi.1 (h x) i, λ h x, differentiable_at_pi.2 (λ i, h i x)⟩ + +-- TODO: find out which version (`φ` or `Φ`) works better with `rw`/`simp` +lemma fderiv_within_pi (h : ∀ i, differentiable_within_at 𝕜 (φ i) s x) + (hs : unique_diff_within_at 𝕜 s x) : + fderiv_within 𝕜 (λ x i, φ i x) s x = pi (λ i, fderiv_within 𝕜 (φ i) s x) := +(has_fderiv_within_at_pi.2 (λ i, (h i).has_fderiv_within_at)).fderiv_within hs + +lemma fderiv_pi (h : ∀ i, differentiable_at 𝕜 (φ i) x) : + fderiv 𝕜 (λ x i, φ i x) x = pi (λ i, fderiv 𝕜 (φ i) x) := +(has_fderiv_at_pi.2 (λ i, (h i).has_fderiv_at)).fderiv + +end pi + +end cartesian_product + +end diff --git a/src/analysis/calculus/fderiv/restrict_scalars.lean b/src/analysis/calculus/fderiv/restrict_scalars.lean new file mode 100644 index 0000000000000..eddc2555b61a3 --- /dev/null +++ b/src/analysis/calculus/fderiv/restrict_scalars.lean @@ -0,0 +1,101 @@ +/- +Copyright (c) 2019 Jeremy Avigad. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jeremy Avigad, Sébastien Gouëzel, Yury Kudryashov +-/ +import analysis.calculus.fderiv.basic + +/-! +# The derivative of the scalar restriction of a linear map + +For detailed documentation of the Fréchet derivative, +see the module docstring of `analysis/calculus/fderiv/basic.lean`. + +This file contains the usual formulas (and existence assertions) for the derivative of +the scalar restriction of a linear map. +-/ + +open filter asymptotics continuous_linear_map set metric +open_locale topology classical nnreal filter asymptotics ennreal + +noncomputable theory + +section restrict_scalars +/-! +### Restricting from `ℂ` to `ℝ`, or generally from `𝕜'` to `𝕜` + +If a function is differentiable over `ℂ`, then it is differentiable over `ℝ`. In this paragraph, +we give variants of this statement, in the general situation where `ℂ` and `ℝ` are replaced +respectively by `𝕜'` and `𝕜` where `𝕜'` is a normed algebra over `𝕜`. +-/ + +variables (𝕜 : Type*) [nontrivially_normed_field 𝕜] +variables {𝕜' : Type*} [nontrivially_normed_field 𝕜'] [normed_algebra 𝕜 𝕜'] +variables {E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] [normed_space 𝕜' E] +variables [is_scalar_tower 𝕜 𝕜' E] +variables {F : Type*} [normed_add_comm_group F] [normed_space 𝕜 F] [normed_space 𝕜' F] +variables [is_scalar_tower 𝕜 𝕜' F] +variables {f : E → F} {f' : E →L[𝕜'] F} {s : set E} {x : E} + +lemma has_strict_fderiv_at.restrict_scalars (h : has_strict_fderiv_at f f' x) : + has_strict_fderiv_at f (f'.restrict_scalars 𝕜) x := h + +lemma has_fderiv_at_filter.restrict_scalars {L} (h : has_fderiv_at_filter f f' x L) : + has_fderiv_at_filter f (f'.restrict_scalars 𝕜) x L := h + +lemma has_fderiv_at.restrict_scalars (h : has_fderiv_at f f' x) : + has_fderiv_at f (f'.restrict_scalars 𝕜) x := h + +lemma has_fderiv_within_at.restrict_scalars (h : has_fderiv_within_at f f' s x) : + has_fderiv_within_at f (f'.restrict_scalars 𝕜) s x := h + +lemma differentiable_at.restrict_scalars (h : differentiable_at 𝕜' f x) : + differentiable_at 𝕜 f x := +(h.has_fderiv_at.restrict_scalars 𝕜).differentiable_at + +lemma differentiable_within_at.restrict_scalars (h : differentiable_within_at 𝕜' f s x) : + differentiable_within_at 𝕜 f s x := +(h.has_fderiv_within_at.restrict_scalars 𝕜).differentiable_within_at + +lemma differentiable_on.restrict_scalars (h : differentiable_on 𝕜' f s) : + differentiable_on 𝕜 f s := +λx hx, (h x hx).restrict_scalars 𝕜 + +lemma differentiable.restrict_scalars (h : differentiable 𝕜' f) : + differentiable 𝕜 f := +λx, (h x).restrict_scalars 𝕜 + +lemma has_fderiv_within_at_of_restrict_scalars + {g' : E →L[𝕜] F} (h : has_fderiv_within_at f g' s x) + (H : f'.restrict_scalars 𝕜 = g') : has_fderiv_within_at f f' s x := +by { rw ← H at h, exact h } + +lemma has_fderiv_at_of_restrict_scalars {g' : E →L[𝕜] F} (h : has_fderiv_at f g' x) + (H : f'.restrict_scalars 𝕜 = g') : has_fderiv_at f f' x := +by { rw ← H at h, exact h } + +lemma differentiable_at.fderiv_restrict_scalars (h : differentiable_at 𝕜' f x) : + fderiv 𝕜 f x = (fderiv 𝕜' f x).restrict_scalars 𝕜 := +(h.has_fderiv_at.restrict_scalars 𝕜).fderiv + +lemma differentiable_within_at_iff_restrict_scalars + (hf : differentiable_within_at 𝕜 f s x) (hs : unique_diff_within_at 𝕜 s x) : + differentiable_within_at 𝕜' f s x ↔ + ∃ (g' : E →L[𝕜'] F), g'.restrict_scalars 𝕜 = fderiv_within 𝕜 f s x := +begin + split, + { rintros ⟨g', hg'⟩, + exact ⟨g', hs.eq (hg'.restrict_scalars 𝕜) hf.has_fderiv_within_at⟩, }, + { rintros ⟨f', hf'⟩, + exact ⟨f', has_fderiv_within_at_of_restrict_scalars 𝕜 hf.has_fderiv_within_at hf'⟩, }, +end + +lemma differentiable_at_iff_restrict_scalars (hf : differentiable_at 𝕜 f x) : + differentiable_at 𝕜' f x ↔ ∃ (g' : E →L[𝕜'] F), g'.restrict_scalars 𝕜 = fderiv 𝕜 f x := +begin + rw [← differentiable_within_at_univ, ← fderiv_within_univ], + exact differentiable_within_at_iff_restrict_scalars 𝕜 + hf.differentiable_within_at unique_diff_within_at_univ, +end + +end restrict_scalars From 7982767093ae38cba236487f9c9dd9cd99f63c16 Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Sun, 14 May 2023 09:06:46 +0000 Subject: [PATCH 1005/1291] chore(analysis/special_functions/gamma): split into 3 files (#19004) This PR splits the existing 1500-line file `analysis.special_functions.gamma` into 3 files. There are no changes to the content of the files (other than top-of-file docstrings and imports). --- src/analysis/special_functions/gamma.lean | 1527 ----------------- .../special_functions/gamma/basic.lean | 673 ++++++++ .../special_functions/gamma/beta.lean | 509 ++++++ .../gamma/bohr_mollerup.lean | 380 ++++ src/analysis/special_functions/gaussian.lean | 2 +- 5 files changed, 1563 insertions(+), 1528 deletions(-) delete mode 100644 src/analysis/special_functions/gamma.lean create mode 100644 src/analysis/special_functions/gamma/basic.lean create mode 100644 src/analysis/special_functions/gamma/beta.lean create mode 100644 src/analysis/special_functions/gamma/bohr_mollerup.lean diff --git a/src/analysis/special_functions/gamma.lean b/src/analysis/special_functions/gamma.lean deleted file mode 100644 index fd7163e6326af..0000000000000 --- a/src/analysis/special_functions/gamma.lean +++ /dev/null @@ -1,1527 +0,0 @@ -/- -Copyright (c) 2022 David Loeffler. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: David Loeffler --/ -import measure_theory.integral.exp_decay -import analysis.special_functions.improper_integrals -import analysis.convolution -import analysis.special_functions.trigonometric.euler_sine_prod - -/-! -# The Gamma and Beta functions - -This file defines the `Γ` function (of a real or complex variable `s`). We define this by Euler's -integral `Γ(s) = ∫ x in Ioi 0, exp (-x) * x ^ (s - 1)` in the range where this integral converges -(i.e., for `0 < s` in the real case, and `0 < re s` in the complex case). - -We show that this integral satisfies `Γ(1) = 1` and `Γ(s + 1) = s * Γ(s)`; hence we can define -`Γ(s)` for all `s` as the unique function satisfying this recurrence and agreeing with Euler's -integral in the convergence range. (If `s = -n` for `n ∈ ℕ`, then the function is undefined, and we -set it to be `0` by convention.) - -## Gamma function: main statements (complex case) - -* `complex.Gamma`: the `Γ` function (of a complex variable). -* `complex.Gamma_eq_integral`: for `0 < re s`, `Γ(s)` agrees with Euler's integral. -* `complex.Gamma_add_one`: for all `s : ℂ` with `s ≠ 0`, we have `Γ (s + 1) = s Γ(s)`. -* `complex.Gamma_nat_eq_factorial`: for all `n : ℕ` we have `Γ (n + 1) = n!`. -* `complex.differentiable_at_Gamma`: `Γ` is complex-differentiable at all `s : ℂ` with - `s ∉ {-n : n ∈ ℕ}`. -* `complex.Gamma_ne_zero`: for all `s : ℂ` with `s ∉ {-n : n ∈ ℕ}` we have `Γ s ≠ 0`. -* `complex.Gamma_seq_tendsto_Gamma`: for all `s`, the limit as `n → ∞` of the sequence - `n ↦ n ^ s * n! / (s * (s + 1) * ... * (s + n))` is `Γ(s)`. -* `complex.Gamma_mul_Gamma_one_sub`: Euler's reflection formula - `Gamma s * Gamma (1 - s) = π / sin π s`. - -## Gamma function: main statements (real case) - -* `real.Gamma`: the `Γ` function (of a real variable). -* Real counterparts of all the properties of the complex Gamma function listed above: - `real.Gamma_eq_integral`, `real.Gamma_add_one`, `real.Gamma_nat_eq_factorial`, - `real.differentiable_at_Gamma`, `real.Gamma_ne_zero`, `real.Gamma_seq_tendsto_Gamma`, - `real.Gamma_mul_Gamma_one_sub`. -* `real.convex_on_log_Gamma` : `log ∘ Γ` is convex on `Ioi 0`. -* `real.eq_Gamma_of_log_convex` : the Bohr-Mollerup theorem, which states that the `Γ` function is - the unique log-convex, positive-valued function on `Ioi 0` satisfying the functional equation - and having `Γ 1 = 1`. - -## Beta function - -* `complex.beta_integral`: the Beta function `Β(u, v)`, where `u`, `v` are complex with positive - real part. -* `complex.Gamma_mul_Gamma_eq_beta_integral`: the formula - `Gamma u * Gamma v = Gamma (u + v) * beta_integral u v`. - -## Tags - -Gamma --/ - -noncomputable theory -open filter interval_integral set real measure_theory asymptotics -open_locale nat topology ennreal big_operators complex_conjugate - -namespace real - -/-- Asymptotic bound for the `Γ` function integrand. -/ -lemma Gamma_integrand_is_o (s : ℝ) : - (λ x:ℝ, exp (-x) * x ^ s) =o[at_top] (λ x:ℝ, exp (-(1/2) * x)) := -begin - refine is_o_of_tendsto (λ x hx, _) _, - { exfalso, exact (exp_pos (-(1 / 2) * x)).ne' hx }, - have : (λ (x:ℝ), exp (-x) * x ^ s / exp (-(1 / 2) * x)) = (λ (x:ℝ), exp ((1 / 2) * x) / x ^ s )⁻¹, - { ext1 x, - field_simp [exp_ne_zero, exp_neg, ← real.exp_add], - left, - ring }, - rw this, - exact (tendsto_exp_mul_div_rpow_at_top s (1 / 2) one_half_pos).inv_tendsto_at_top, -end - -/-- The Euler integral for the `Γ` function converges for positive real `s`. -/ -lemma Gamma_integral_convergent {s : ℝ} (h : 0 < s) : - integrable_on (λ x:ℝ, exp (-x) * x ^ (s - 1)) (Ioi 0) := -begin - rw [←Ioc_union_Ioi_eq_Ioi (@zero_le_one ℝ _ _ _ _), integrable_on_union], - split, - { rw ←integrable_on_Icc_iff_integrable_on_Ioc, - refine integrable_on.continuous_on_mul continuous_on_id.neg.exp _ is_compact_Icc, - refine (interval_integrable_iff_integrable_Icc_of_le zero_le_one).mp _, - exact interval_integrable_rpow' (by linarith), }, - { refine integrable_of_is_O_exp_neg one_half_pos _ (Gamma_integrand_is_o _ ).is_O, - refine continuous_on_id.neg.exp.mul (continuous_on_id.rpow_const _), - intros x hx, - exact or.inl ((zero_lt_one : (0 : ℝ) < 1).trans_le hx).ne' } -end - -end real - -namespace complex -/- Technical note: In defining the Gamma integrand exp (-x) * x ^ (s - 1) for s complex, we have to -make a choice between ↑(real.exp (-x)), complex.exp (↑(-x)), and complex.exp (-↑x), all of which are -equal but not definitionally so. We use the first of these throughout. -/ - - -/-- The integral defining the `Γ` function converges for complex `s` with `0 < re s`. - -This is proved by reduction to the real case. -/ -lemma Gamma_integral_convergent {s : ℂ} (hs : 0 < s.re) : - integrable_on (λ x, (-x).exp * x ^ (s - 1) : ℝ → ℂ) (Ioi 0) := -begin - split, - { refine continuous_on.ae_strongly_measurable _ measurable_set_Ioi, - apply (continuous_of_real.comp continuous_neg.exp).continuous_on.mul, - apply continuous_at.continuous_on, - intros x hx, - have : continuous_at (λ x:ℂ, x ^ (s - 1)) ↑x, - { apply continuous_at_cpow_const, rw of_real_re, exact or.inl hx, }, - exact continuous_at.comp this continuous_of_real.continuous_at }, - { rw ←has_finite_integral_norm_iff, - refine has_finite_integral.congr (real.Gamma_integral_convergent hs).2 _, - refine (ae_restrict_iff' measurable_set_Ioi).mpr (ae_of_all _ (λ x hx, _)), - dsimp only, - rw [norm_eq_abs, map_mul, abs_of_nonneg $ le_of_lt $ exp_pos $ -x, - abs_cpow_eq_rpow_re_of_pos hx _], - simp } -end - -/-- Euler's integral for the `Γ` function (of a complex variable `s`), defined as -`∫ x in Ioi 0, exp (-x) * x ^ (s - 1)`. - -See `complex.Gamma_integral_convergent` for a proof of the convergence of the integral for -`0 < re s`. -/ -def Gamma_integral (s : ℂ) : ℂ := ∫ x in Ioi (0:ℝ), ↑(-x).exp * ↑x ^ (s - 1) - -lemma Gamma_integral_conj (s : ℂ) : Gamma_integral (conj s) = conj (Gamma_integral s) := -begin - rw [Gamma_integral, Gamma_integral, ←integral_conj], - refine set_integral_congr measurable_set_Ioi (λ x hx, _), - dsimp only, - rw [ring_hom.map_mul, conj_of_real, cpow_def_of_ne_zero (of_real_ne_zero.mpr (ne_of_gt hx)), - cpow_def_of_ne_zero (of_real_ne_zero.mpr (ne_of_gt hx)), ←exp_conj, ring_hom.map_mul, - ←of_real_log (le_of_lt hx), conj_of_real, ring_hom.map_sub, ring_hom.map_one], -end - -lemma Gamma_integral_of_real (s : ℝ) : - Gamma_integral ↑s = ↑(∫ x:ℝ in Ioi 0, real.exp (-x) * x ^ (s - 1)) := -begin - rw [Gamma_integral, ←_root_.integral_of_real], - refine set_integral_congr measurable_set_Ioi _, - intros x hx, dsimp only, - rw [of_real_mul, of_real_cpow (mem_Ioi.mp hx).le], - simp, -end - -lemma Gamma_integral_one : Gamma_integral 1 = 1 := -by simpa only [←of_real_one, Gamma_integral_of_real, of_real_inj, sub_self, - rpow_zero, mul_one] using integral_exp_neg_Ioi_zero - -end complex - -/-! Now we establish the recurrence relation `Γ(s + 1) = s * Γ(s)` using integration by parts. -/ - -namespace complex - -section Gamma_recurrence - -/-- The indefinite version of the `Γ` function, `Γ(s, X) = ∫ x ∈ 0..X, exp(-x) x ^ (s - 1)`. -/ -def partial_Gamma (s : ℂ) (X : ℝ) : ℂ := ∫ x in 0..X, (-x).exp * x ^ (s - 1) - -lemma tendsto_partial_Gamma {s : ℂ} (hs: 0 < s.re) : - tendsto (λ X:ℝ, partial_Gamma s X) at_top (𝓝 $ Gamma_integral s) := -interval_integral_tendsto_integral_Ioi 0 (Gamma_integral_convergent hs) tendsto_id - -private lemma Gamma_integrand_interval_integrable (s : ℂ) {X : ℝ} (hs : 0 < s.re) (hX : 0 ≤ X): - interval_integrable (λ x, (-x).exp * x ^ (s - 1) : ℝ → ℂ) volume 0 X := -begin - rw interval_integrable_iff_integrable_Ioc_of_le hX, - exact integrable_on.mono_set (Gamma_integral_convergent hs) Ioc_subset_Ioi_self -end - -private lemma Gamma_integrand_deriv_integrable_A {s : ℂ} (hs : 0 < s.re) {X : ℝ} (hX : 0 ≤ X): - interval_integrable (λ x, -((-x).exp * x ^ s) : ℝ → ℂ) volume 0 X := -begin - convert (Gamma_integrand_interval_integrable (s+1) _ hX).neg, - { ext1, simp only [add_sub_cancel, pi.neg_apply] }, - { simp only [add_re, one_re], linarith,}, -end - -private lemma Gamma_integrand_deriv_integrable_B {s : ℂ} (hs : 0 < s.re) {Y : ℝ} (hY : 0 ≤ Y) : - interval_integrable (λ (x : ℝ), (-x).exp * (s * x ^ (s - 1)) : ℝ → ℂ) volume 0 Y := -begin - have : (λ x, (-x).exp * (s * x ^ (s - 1)) : ℝ → ℂ) = - (λ x, s * ((-x).exp * x ^ (s - 1)) : ℝ → ℂ), - { ext1, ring, }, - rw [this, interval_integrable_iff_integrable_Ioc_of_le hY], - split, - { refine (continuous_on_const.mul _).ae_strongly_measurable measurable_set_Ioc, - apply (continuous_of_real.comp continuous_neg.exp).continuous_on.mul, - apply continuous_at.continuous_on, - intros x hx, - refine (_ : continuous_at (λ x:ℂ, x ^ (s - 1)) _).comp continuous_of_real.continuous_at, - apply continuous_at_cpow_const, rw of_real_re, exact or.inl hx.1, }, - rw ←has_finite_integral_norm_iff, - simp_rw [norm_eq_abs, map_mul], - refine (((real.Gamma_integral_convergent hs).mono_set - Ioc_subset_Ioi_self).has_finite_integral.congr _).const_mul _, - rw [eventually_eq, ae_restrict_iff'], - { apply ae_of_all, intros x hx, - rw [abs_of_nonneg (exp_pos _).le,abs_cpow_eq_rpow_re_of_pos hx.1], - simp }, - { exact measurable_set_Ioc}, -end - -/-- The recurrence relation for the indefinite version of the `Γ` function. -/ -lemma partial_Gamma_add_one {s : ℂ} (hs: 0 < s.re) {X : ℝ} (hX : 0 ≤ X) : - partial_Gamma (s + 1) X = s * partial_Gamma s X - (-X).exp * X ^ s := -begin - rw [partial_Gamma, partial_Gamma, add_sub_cancel], - have F_der_I: (∀ (x:ℝ), (x ∈ Ioo 0 X) → has_deriv_at (λ x, (-x).exp * x ^ s : ℝ → ℂ) - ( -((-x).exp * x ^ s) + (-x).exp * (s * x ^ (s - 1))) x), - { intros x hx, - have d1 : has_deriv_at (λ (y: ℝ), (-y).exp) (-(-x).exp) x, - { simpa using (has_deriv_at_neg x).exp }, - have d2 : has_deriv_at (λ (y : ℝ), ↑y ^ s) (s * x ^ (s - 1)) x, - { have t := @has_deriv_at.cpow_const _ _ _ s (has_deriv_at_id ↑x) _, - simpa only [mul_one] using t.comp_of_real, - simpa only [id.def, of_real_re, of_real_im, - ne.def, eq_self_iff_true, not_true, or_false, mul_one] using hx.1, }, - simpa only [of_real_neg, neg_mul] using d1.of_real_comp.mul d2 }, - have cont := (continuous_of_real.comp continuous_neg.exp).mul - (continuous_of_real_cpow_const hs), - have der_ible := (Gamma_integrand_deriv_integrable_A hs hX).add - (Gamma_integrand_deriv_integrable_B hs hX), - have int_eval := integral_eq_sub_of_has_deriv_at_of_le hX cont.continuous_on F_der_I der_ible, - -- We are basically done here but manipulating the output into the right form is fiddly. - apply_fun (λ x:ℂ, -x) at int_eval, - rw [interval_integral.integral_add (Gamma_integrand_deriv_integrable_A hs hX) - (Gamma_integrand_deriv_integrable_B hs hX), interval_integral.integral_neg, neg_add, neg_neg] - at int_eval, - rw [eq_sub_of_add_eq int_eval, sub_neg_eq_add, neg_sub, add_comm, add_sub], - simp only [sub_left_inj, add_left_inj], - have : (λ x, (-x).exp * (s * x ^ (s - 1)) : ℝ → ℂ) = (λ x, s * (-x).exp * x ^ (s - 1) : ℝ → ℂ), - { ext1, ring,}, - rw this, - have t := @integral_const_mul 0 X volume _ _ s (λ x:ℝ, (-x).exp * x ^ (s - 1)), - dsimp at t, rw [←t, of_real_zero, zero_cpow], - { rw [mul_zero, add_zero], congr', ext1, ring }, - { contrapose! hs, rw [hs, zero_re] } -end - -/-- The recurrence relation for the `Γ` integral. -/ -theorem Gamma_integral_add_one {s : ℂ} (hs: 0 < s.re) : - Gamma_integral (s + 1) = s * Gamma_integral s := -begin - suffices : tendsto (s+1).partial_Gamma at_top (𝓝 $ s * Gamma_integral s), - { refine tendsto_nhds_unique _ this, - apply tendsto_partial_Gamma, rw [add_re, one_re], linarith, }, - have : (λ X:ℝ, s * partial_Gamma s X - X ^ s * (-X).exp) =ᶠ[at_top] (s+1).partial_Gamma, - { apply eventually_eq_of_mem (Ici_mem_at_top (0:ℝ)), - intros X hX, - rw partial_Gamma_add_one hs (mem_Ici.mp hX), - ring_nf, }, - refine tendsto.congr' this _, - suffices : tendsto (λ X, -X ^ s * (-X).exp : ℝ → ℂ) at_top (𝓝 0), - { simpa using tendsto.add (tendsto.const_mul s (tendsto_partial_Gamma hs)) this }, - rw tendsto_zero_iff_norm_tendsto_zero, - have : (λ (e : ℝ), ‖-(e:ℂ) ^ s * (-e).exp‖ ) =ᶠ[at_top] (λ (e : ℝ), e ^ s.re * (-1 * e).exp ), - { refine eventually_eq_of_mem (Ioi_mem_at_top 0) _, - intros x hx, dsimp only, - rw [norm_eq_abs, map_mul, abs.map_neg, abs_cpow_eq_rpow_re_of_pos hx, - abs_of_nonneg (exp_pos(-x)).le, neg_mul, one_mul],}, - exact (tendsto_congr' this).mpr (tendsto_rpow_mul_exp_neg_mul_at_top_nhds_0 _ _ zero_lt_one), -end - -end Gamma_recurrence - -/-! Now we define `Γ(s)` on the whole complex plane, by recursion. -/ - -section Gamma_def - -/-- The `n`th function in this family is `Γ(s)` if `-n < s.re`, and junk otherwise. -/ -noncomputable def Gamma_aux : ℕ → (ℂ → ℂ) -| 0 := Gamma_integral -| (n+1) := λ s:ℂ, (Gamma_aux n (s+1)) / s - -lemma Gamma_aux_recurrence1 (s : ℂ) (n : ℕ) (h1 : -s.re < ↑n) : - Gamma_aux n s = Gamma_aux n (s+1) / s := -begin - induction n with n hn generalizing s, - { simp only [nat.cast_zero, neg_lt_zero] at h1, - dsimp only [Gamma_aux], rw Gamma_integral_add_one h1, - rw [mul_comm, mul_div_cancel], contrapose! h1, rw h1, - simp }, - { dsimp only [Gamma_aux], - have hh1 : -(s+1).re < n, - { rw [nat.succ_eq_add_one, nat.cast_add, nat.cast_one] at h1, - rw [add_re, one_re], linarith, }, - rw ←(hn (s+1) hh1) } -end - -lemma Gamma_aux_recurrence2 (s : ℂ) (n : ℕ) (h1 : -s.re < ↑n) : - Gamma_aux n s = Gamma_aux (n+1) s := -begin - cases n, - { simp only [nat.cast_zero, neg_lt_zero] at h1, - dsimp only [Gamma_aux], - rw [Gamma_integral_add_one h1, mul_div_cancel_left], - rintro rfl, - rw [zero_re] at h1, - exact h1.false }, - { dsimp only [Gamma_aux], - have : (Gamma_aux n (s + 1 + 1)) / (s+1) = Gamma_aux n (s + 1), - { have hh1 : -(s+1).re < n, - { rw [nat.succ_eq_add_one, nat.cast_add, nat.cast_one] at h1, - rw [add_re, one_re], linarith, }, - rw Gamma_aux_recurrence1 (s+1) n hh1, }, - rw this }, -end - -/-- The `Γ` function (of a complex variable `s`). -/ -@[pp_nodot] def Gamma (s : ℂ) : ℂ := Gamma_aux ⌊1 - s.re⌋₊ s - -lemma Gamma_eq_Gamma_aux (s : ℂ) (n : ℕ) (h1 : -s.re < ↑n) : Gamma s = Gamma_aux n s := -begin - have u : ∀ (k : ℕ), Gamma_aux (⌊1 - s.re⌋₊ + k) s = Gamma s, - { intro k, induction k with k hk, - { simp [Gamma],}, - { rw [←hk, nat.succ_eq_add_one, ←add_assoc], - refine (Gamma_aux_recurrence2 s (⌊1 - s.re⌋₊ + k) _).symm, - rw nat.cast_add, - have i0 := nat.sub_one_lt_floor (1 - s.re), - simp only [sub_sub_cancel_left] at i0, - refine lt_add_of_lt_of_nonneg i0 _, - rw [←nat.cast_zero, nat.cast_le], exact nat.zero_le k, } }, - convert (u $ n - ⌊1 - s.re⌋₊).symm, rw nat.add_sub_of_le, - by_cases (0 ≤ 1 - s.re), - { apply nat.le_of_lt_succ, - exact_mod_cast lt_of_le_of_lt (nat.floor_le h) (by linarith : 1 - s.re < n + 1) }, - { rw nat.floor_of_nonpos, linarith, linarith }, -end - -/-- The recurrence relation for the `Γ` function. -/ -theorem Gamma_add_one (s : ℂ) (h2 : s ≠ 0) : Gamma (s+1) = s * Gamma s := -begin - let n := ⌊1 - s.re⌋₊, - have t1 : -s.re < n, - { simpa only [sub_sub_cancel_left] using nat.sub_one_lt_floor (1 - s.re) }, - have t2 : -(s+1).re < n, - { rw [add_re, one_re], linarith, }, - rw [Gamma_eq_Gamma_aux s n t1, Gamma_eq_Gamma_aux (s+1) n t2, Gamma_aux_recurrence1 s n t1], - field_simp, ring, -end - -theorem Gamma_eq_integral {s : ℂ} (hs : 0 < s.re) : Gamma s = Gamma_integral s := -Gamma_eq_Gamma_aux s 0 (by { norm_cast, linarith }) - -lemma Gamma_one : Gamma 1 = 1 := -by { rw Gamma_eq_integral, simpa using Gamma_integral_one, simp } - -theorem Gamma_nat_eq_factorial (n : ℕ) : Gamma (n+1) = n! := -begin - induction n with n hn, - { simpa using Gamma_one }, - { rw (Gamma_add_one n.succ $ nat.cast_ne_zero.mpr $ nat.succ_ne_zero n), - simp only [nat.cast_succ, nat.factorial_succ, nat.cast_mul], congr, exact hn }, -end - -/-- At `0` the Gamma function is undefined; by convention we assign it the value `0`. -/ -lemma Gamma_zero : Gamma 0 = 0 := -by simp_rw [Gamma, zero_re, sub_zero, nat.floor_one, Gamma_aux, div_zero] - -/-- At `-n` for `n ∈ ℕ`, the Gamma function is undefined; by convention we assign it the value 0. -/ -lemma Gamma_neg_nat_eq_zero (n : ℕ) : Gamma (-n) = 0 := -begin - induction n with n IH, - { rw [nat.cast_zero, neg_zero, Gamma_zero] }, - { have A : -(n.succ : ℂ) ≠ 0, - { rw [neg_ne_zero, nat.cast_ne_zero], - apply nat.succ_ne_zero }, - have : -(n:ℂ) = -↑n.succ + 1, by simp, - rw [this, Gamma_add_one _ A] at IH, - contrapose! IH, - exact mul_ne_zero A IH } -end - -lemma Gamma_conj (s : ℂ) : Gamma (conj s) = conj (Gamma s) := -begin - suffices : ∀ (n:ℕ) (s:ℂ) , Gamma_aux n (conj s) = conj (Gamma_aux n s), from this _ _, - intro n, - induction n with n IH, - { rw Gamma_aux, exact Gamma_integral_conj, }, - { intro s, - rw Gamma_aux, - dsimp only, - rw [div_eq_mul_inv _ s, ring_hom.map_mul, conj_inv, ←div_eq_mul_inv], - suffices : conj s + 1 = conj (s + 1), by rw [this, IH], - rw [ring_hom.map_add, ring_hom.map_one] } -end - -end Gamma_def - -end complex - -/-! Now check that the `Γ` function is differentiable, wherever this makes sense. -/ - -section Gamma_has_deriv - -/-- Integrand for the derivative of the `Γ` function -/ -def dGamma_integrand (s : ℂ) (x : ℝ) : ℂ := exp (-x) * log x * x ^ (s - 1) - -/-- Integrand for the absolute value of the derivative of the `Γ` function -/ -def dGamma_integrand_real (s x : ℝ) : ℝ := |exp (-x) * log x * x ^ (s - 1)| - -lemma dGamma_integrand_is_o_at_top (s : ℝ) : - (λ x : ℝ, exp (-x) * log x * x ^ (s - 1)) =o[at_top] (λ x, exp (-(1/2) * x)) := -begin - refine is_o_of_tendsto (λ x hx, _) _, - { exfalso, exact (-(1/2) * x).exp_pos.ne' hx, }, - have : eventually_eq at_top (λ (x : ℝ), exp (-x) * log x * x ^ (s - 1) / exp (-(1 / 2) * x)) - (λ (x : ℝ), (λ z:ℝ, exp (1 / 2 * z) / z ^ s) x * (λ z:ℝ, z / log z) x)⁻¹, - { refine eventually_of_mem (Ioi_mem_at_top 1) _, - intros x hx, dsimp, - replace hx := lt_trans zero_lt_one (mem_Ioi.mp hx), - rw [real.exp_neg, neg_mul, real.exp_neg, rpow_sub hx], - have : exp x = exp(x/2) * exp(x/2), - { rw [←real.exp_add, add_halves], }, - rw this, field_simp [hx.ne', exp_ne_zero (x/2)], ring, }, - refine tendsto.congr' this.symm (tendsto.inv_tendsto_at_top _), - apply tendsto.at_top_mul_at_top (tendsto_exp_mul_div_rpow_at_top s (1/2) one_half_pos), - refine tendsto.congr' _ ((tendsto_exp_div_pow_at_top 1).comp tendsto_log_at_top), - apply eventually_eq_of_mem (Ioi_mem_at_top (0:ℝ)), - intros x hx, simp [exp_log hx], -end - -/-- Absolute convergence of the integral which will give the derivative of the `Γ` function on -`1 < re s`. -/ -lemma dGamma_integral_abs_convergent (s : ℝ) (hs : 1 < s) : - integrable_on (λ x:ℝ, ‖exp (-x) * log x * x ^ (s-1)‖) (Ioi 0) := -begin - rw [←Ioc_union_Ioi_eq_Ioi (@zero_le_one ℝ _ _ _ _), integrable_on_union], - refine ⟨⟨_, _⟩, _⟩, - { refine continuous_on.ae_strongly_measurable (continuous_on.mul _ _).norm measurable_set_Ioc, - { refine (continuous_exp.comp continuous_neg).continuous_on.mul (continuous_on_log.mono _), - simp, }, - { apply continuous_on_id.rpow_const, intros x hx, right, linarith }, }, - { apply has_finite_integral_of_bounded, - swap, { exact 1 / (s - 1), }, - refine (ae_restrict_iff' measurable_set_Ioc).mpr (ae_of_all _ (λ x hx, _)), - rw [norm_norm, norm_eq_abs, mul_assoc, abs_mul, ←one_mul (1 / (s - 1))], - refine mul_le_mul _ _ (abs_nonneg _) zero_le_one, - { rw [abs_of_pos (exp_pos(-x)), exp_le_one_iff, neg_le, neg_zero], exact hx.1.le }, - { exact (abs_log_mul_self_rpow_lt x (s-1) hx.1 hx.2 (sub_pos.mpr hs)).le }, }, - { have := (dGamma_integrand_is_o_at_top s).is_O.norm_left, - refine integrable_of_is_O_exp_neg one_half_pos (continuous_on.mul _ _).norm this, - { refine (continuous_exp.comp continuous_neg).continuous_on.mul (continuous_on_log.mono _), - simp, }, - { apply continuous_at.continuous_on (λ x hx, _), - apply continuous_at_id.rpow continuous_at_const, - dsimp, right, linarith, }, } -end - -/-- A uniform bound for the `s`-derivative of the `Γ` integrand for `s` in vertical strips. -/ -lemma loc_unif_bound_dGamma_integrand {t : ℂ} {s1 s2 x : ℝ} (ht1 : s1 ≤ t.re) - (ht2: t.re ≤ s2) (hx : 0 < x) : - ‖dGamma_integrand t x‖ ≤ dGamma_integrand_real s1 x + dGamma_integrand_real s2 x := -begin - rcases le_or_lt 1 x with h|h, - { -- case 1 ≤ x - refine le_add_of_nonneg_of_le (abs_nonneg _) _, - rw [dGamma_integrand, dGamma_integrand_real, complex.norm_eq_abs, map_mul, abs_mul, - ←complex.of_real_mul, complex.abs_of_real], - refine mul_le_mul_of_nonneg_left _ (abs_nonneg _), - rw complex.abs_cpow_eq_rpow_re_of_pos hx, - refine le_trans _ (le_abs_self _), - apply rpow_le_rpow_of_exponent_le h, - rw [complex.sub_re, complex.one_re], linarith, }, - { refine le_add_of_le_of_nonneg _ (abs_nonneg _), - rw [dGamma_integrand, dGamma_integrand_real, complex.norm_eq_abs, map_mul, abs_mul, - ←complex.of_real_mul, complex.abs_of_real], - refine mul_le_mul_of_nonneg_left _ (abs_nonneg _), - rw complex.abs_cpow_eq_rpow_re_of_pos hx, - refine le_trans _ (le_abs_self _), - apply rpow_le_rpow_of_exponent_ge hx h.le, - rw [complex.sub_re, complex.one_re], linarith, }, -end - -namespace complex - -/-- The derivative of the `Γ` integral, at any `s ∈ ℂ` with `1 < re s`, is given by the integral -of `exp (-x) * log x * x ^ (s - 1)` over `[0, ∞)`. -/ -theorem has_deriv_at_Gamma_integral {s : ℂ} (hs : 1 < s.re) : - (integrable_on (λ x, real.exp (-x) * real.log x * x ^ (s - 1) : ℝ → ℂ) (Ioi 0) volume) ∧ - (has_deriv_at Gamma_integral (∫ x:ℝ in Ioi 0, real.exp (-x) * real.log x * x ^ (s - 1)) s) := -begin - let ε := (s.re - 1) / 2, - let μ := volume.restrict (Ioi (0:ℝ)), - let bound := (λ x:ℝ, dGamma_integrand_real (s.re - ε) x + dGamma_integrand_real (s.re + ε) x), - have cont : ∀ (t : ℂ), continuous_on (λ x, real.exp (-x) * x ^ (t - 1) : ℝ → ℂ) (Ioi 0), - { intro t, apply (continuous_of_real.comp continuous_neg.exp).continuous_on.mul, - apply continuous_at.continuous_on, intros x hx, - refine (continuous_at_cpow_const _).comp continuous_of_real.continuous_at, - exact or.inl hx, }, - have eps_pos: 0 < ε := div_pos (sub_pos.mpr hs) zero_lt_two, - have hF_meas : ∀ᶠ (t : ℂ) in 𝓝 s, - ae_strongly_measurable (λ x, real.exp(-x) * x ^ (t - 1) : ℝ → ℂ) μ, - { apply eventually_of_forall, intro t, - exact (cont t).ae_strongly_measurable measurable_set_Ioi, }, - have hF'_meas : ae_strongly_measurable (dGamma_integrand s) μ, - { refine continuous_on.ae_strongly_measurable _ measurable_set_Ioi, - have : dGamma_integrand s = (λ x, real.exp (-x) * x ^ (s - 1) * real.log x : ℝ → ℂ), - { ext1, simp only [dGamma_integrand], ring }, - rw this, - refine continuous_on.mul (cont s) (continuous_at.continuous_on _), - exact λ x hx, continuous_of_real.continuous_at.comp (continuous_at_log (mem_Ioi.mp hx).ne'), }, - have h_bound : ∀ᵐ (x : ℝ) ∂μ, ∀ (t : ℂ), t ∈ metric.ball s ε → ‖ dGamma_integrand t x ‖ ≤ bound x, - { refine (ae_restrict_iff' measurable_set_Ioi).mpr (ae_of_all _ (λ x hx, _)), - intros t ht, - rw [metric.mem_ball, complex.dist_eq] at ht, - replace ht := lt_of_le_of_lt (complex.abs_re_le_abs $ t - s ) ht, - rw [complex.sub_re, @abs_sub_lt_iff ℝ _ t.re s.re ((s.re - 1) / 2) ] at ht, - refine loc_unif_bound_dGamma_integrand _ _ hx, - all_goals { simp only [ε], linarith } }, - have bound_integrable : integrable bound μ, - { apply integrable.add, - { refine dGamma_integral_abs_convergent (s.re - ε) _, - field_simp, rw one_lt_div, - { linarith }, { exact zero_lt_two }, }, - { refine dGamma_integral_abs_convergent (s.re + ε) _, linarith, }, }, - have h_diff : ∀ᵐ (x : ℝ) ∂μ, ∀ (t : ℂ), t ∈ metric.ball s ε - → has_deriv_at (λ u, real.exp (-x) * x ^ (u - 1) : ℂ → ℂ) (dGamma_integrand t x) t, - { refine (ae_restrict_iff' measurable_set_Ioi).mpr (ae_of_all _ (λ x hx, _)), - intros t ht, rw mem_Ioi at hx, - simp only [dGamma_integrand], - rw mul_assoc, - apply has_deriv_at.const_mul, - rw [of_real_log hx.le, mul_comm], - have := ((has_deriv_at_id t).sub_const 1).const_cpow (or.inl (of_real_ne_zero.mpr hx.ne')), - rwa mul_one at this }, - exact (has_deriv_at_integral_of_dominated_loc_of_deriv_le eps_pos hF_meas - (Gamma_integral_convergent (zero_lt_one.trans hs)) hF'_meas h_bound bound_integrable h_diff), -end - -lemma differentiable_at_Gamma_aux (s : ℂ) (n : ℕ) (h1 : (1 - s.re) < n ) (h2 : ∀ m : ℕ, s ≠ -m) : - differentiable_at ℂ (Gamma_aux n) s := -begin - induction n with n hn generalizing s, - { refine (has_deriv_at_Gamma_integral _).2.differentiable_at, - rw nat.cast_zero at h1, linarith }, - { dsimp only [Gamma_aux], - specialize hn (s + 1), - have a : 1 - (s + 1).re < ↑n, - { rw nat.cast_succ at h1, rw [complex.add_re, complex.one_re], linarith }, - have b : ∀ m : ℕ, s + 1 ≠ -m, - { intro m, have := h2 (1 + m), - contrapose! this, - rw ←eq_sub_iff_add_eq at this, - simpa using this }, - refine differentiable_at.div (differentiable_at.comp _ (hn a b) _) _ _, - simp, simp, simpa using h2 0 } -end - -theorem differentiable_at_Gamma (s : ℂ) (hs : ∀ m : ℕ, s ≠ -m) : differentiable_at ℂ Gamma s := -begin - let n := ⌊1 - s.re⌋₊ + 1, - have hn : 1 - s.re < n := by exact_mod_cast nat.lt_floor_add_one (1 - s.re), - apply (differentiable_at_Gamma_aux s n hn hs).congr_of_eventually_eq, - let S := { t : ℂ | 1 - t.re < n }, - have : S ∈ 𝓝 s, - { rw mem_nhds_iff, use S, - refine ⟨subset.rfl, _, hn⟩, - have : S = re⁻¹' Ioi (1 - n : ℝ), - { ext, rw [preimage,Ioi, mem_set_of_eq, mem_set_of_eq, mem_set_of_eq], exact sub_lt_comm }, - rw this, - refine continuous.is_open_preimage continuous_re _ is_open_Ioi, }, - apply eventually_eq_of_mem this, - intros t ht, rw mem_set_of_eq at ht, - apply Gamma_eq_Gamma_aux, linarith, -end - -end complex - -end Gamma_has_deriv - -namespace real - -/-- The `Γ` function (of a real variable `s`). -/ -@[pp_nodot] def Gamma (s : ℝ) : ℝ := (complex.Gamma s).re - -lemma Gamma_eq_integral {s : ℝ} (hs : 0 < s) : Gamma s = ∫ x in Ioi 0, exp (-x) * x ^ (s - 1) := -begin - rw [Gamma, complex.Gamma_eq_integral (by rwa complex.of_real_re : 0 < complex.re s)], - dsimp only [complex.Gamma_integral], - simp_rw [←complex.of_real_one, ←complex.of_real_sub], - suffices : ∫ (x : ℝ) in Ioi 0, ↑(exp (-x)) * (x : ℂ) ^ ((s - 1 : ℝ) : ℂ) = - ∫ (x : ℝ) in Ioi 0, ((exp (-x) * x ^ (s - 1) : ℝ) : ℂ), - { rw [this, _root_.integral_of_real, complex.of_real_re], }, - refine set_integral_congr measurable_set_Ioi (λ x hx, _), - push_cast, - rw complex.of_real_cpow (le_of_lt hx), - push_cast, -end - -lemma Gamma_add_one {s : ℝ} (hs : s ≠ 0) : Gamma (s + 1) = s * Gamma s := -begin - simp_rw Gamma, - rw [complex.of_real_add, complex.of_real_one, complex.Gamma_add_one, complex.of_real_mul_re], - rwa complex.of_real_ne_zero, -end - -lemma Gamma_one : Gamma 1 = 1 := -by rw [Gamma, complex.of_real_one, complex.Gamma_one, complex.one_re] - -lemma _root_.complex.Gamma_of_real (s : ℝ) : complex.Gamma (s : ℂ) = Gamma s := -by rw [Gamma, eq_comm, ←complex.conj_eq_iff_re, ←complex.Gamma_conj, complex.conj_of_real] - -theorem Gamma_nat_eq_factorial (n : ℕ) : Gamma (n + 1) = n! := -by rw [Gamma, complex.of_real_add, complex.of_real_nat_cast, complex.of_real_one, - complex.Gamma_nat_eq_factorial, ←complex.of_real_nat_cast, complex.of_real_re] - -/-- At `0` the Gamma function is undefined; by convention we assign it the value `0`. -/ -lemma Gamma_zero : Gamma 0 = 0 := -by simpa only [←complex.of_real_zero, complex.Gamma_of_real, complex.of_real_inj] - using complex.Gamma_zero - -/-- At `-n` for `n ∈ ℕ`, the Gamma function is undefined; by convention we assign it the value `0`. --/ -lemma Gamma_neg_nat_eq_zero (n : ℕ) : Gamma (-n) = 0 := -begin - simpa only [←complex.of_real_nat_cast, ←complex.of_real_neg, complex.Gamma_of_real, - complex.of_real_eq_zero] using complex.Gamma_neg_nat_eq_zero n, -end - -lemma Gamma_pos_of_pos {s : ℝ} (hs : 0 < s) : 0 < Gamma s := -begin - rw Gamma_eq_integral hs, - have : function.support (λ (x : ℝ), exp (-x) * x ^ (s - 1)) ∩ Ioi 0 = Ioi 0, - { rw inter_eq_right_iff_subset, - intros x hx, - rw function.mem_support, - exact mul_ne_zero (exp_pos _).ne' (rpow_pos_of_pos hx _).ne' }, - rw set_integral_pos_iff_support_of_nonneg_ae, - { rw [this, volume_Ioi, ←ennreal.of_real_zero], - exact ennreal.of_real_lt_top }, - { refine eventually_of_mem (self_mem_ae_restrict measurable_set_Ioi) _, - exact λ x hx, (mul_pos (exp_pos _) (rpow_pos_of_pos hx _)).le }, - { exact Gamma_integral_convergent hs }, -end - -/-- The Gamma function does not vanish on `ℝ` (except at non-positive integers, where the function -is mathematically undefined and we set it to `0` by convention). -/ -lemma Gamma_ne_zero {s : ℝ} (hs : ∀ m : ℕ, s ≠ -m) : Gamma s ≠ 0 := -begin - suffices : ∀ {n : ℕ}, (-(n:ℝ) < s) → Gamma s ≠ 0, - { apply this, - swap, use (⌊-s⌋₊ + 1), - rw [neg_lt, nat.cast_add, nat.cast_one], - exact nat.lt_floor_add_one _ }, - intro n, - induction n generalizing s, - { intro hs, - refine (Gamma_pos_of_pos _).ne', - rwa [nat.cast_zero, neg_zero] at hs }, - { intro hs', - have : Gamma (s + 1) ≠ 0, - { apply n_ih, - { intro m, - specialize hs (1 + m), - contrapose! hs, - rw ←eq_sub_iff_add_eq at hs, - rw hs, - push_cast, - ring }, - { rw [nat.succ_eq_add_one, nat.cast_add, nat.cast_one, neg_add] at hs', - linarith } }, - rw [Gamma_add_one, mul_ne_zero_iff] at this, - { exact this.2 }, - { simpa using hs 0 } }, -end - -lemma Gamma_eq_zero_iff (s : ℝ) : Gamma s = 0 ↔ ∃ m : ℕ, s = -m := -⟨by { contrapose!, exact Gamma_ne_zero }, by { rintro ⟨m, rfl⟩, exact Gamma_neg_nat_eq_zero m }⟩ - -lemma differentiable_at_Gamma {s : ℝ} (hs : ∀ m : ℕ, s ≠ -m) : differentiable_at ℝ Gamma s := -begin - refine ((complex.differentiable_at_Gamma _ _).has_deriv_at).real_of_complex.differentiable_at, - simp_rw [←complex.of_real_nat_cast, ←complex.of_real_neg, ne.def, complex.of_real_inj], - exact hs, -end - -/-- Log-convexity of the Gamma function on the positive reals (stated in multiplicative form), -proved using the Hölder inequality applied to Euler's integral. -/ -lemma Gamma_mul_add_mul_le_rpow_Gamma_mul_rpow_Gamma {s t a b : ℝ} - (hs : 0 < s) (ht : 0 < t) (ha : 0 < a) (hb : 0 < b) (hab : a + b = 1) : - Gamma (a * s + b * t) ≤ Gamma s ^ a * Gamma t ^ b := -begin - -- We will apply Hölder's inequality, for the conjugate exponents `p = 1 / a` - -- and `q = 1 / b`, to the functions `f a s` and `f b t`, where `f` is as follows: - let f : ℝ → ℝ → ℝ → ℝ := λ c u x, exp (-c * x) * x ^ (c * (u - 1)), - have e : is_conjugate_exponent (1 / a) (1 / b) := real.is_conjugate_exponent_one_div ha hb hab, - have hab' : b = 1 - a := by linarith, - have hst : 0 < a * s + b * t := add_pos (mul_pos ha hs) (mul_pos hb ht), - -- some properties of f: - have posf : ∀ (c u x : ℝ), x ∈ Ioi (0:ℝ) → 0 ≤ f c u x := - λ c u x hx, mul_nonneg (exp_pos _).le (rpow_pos_of_pos hx _).le, - have posf' : ∀ (c u : ℝ), ∀ᵐ (x : ℝ) ∂volume.restrict (Ioi 0), 0 ≤ f c u x := - λ c u, (ae_restrict_iff' measurable_set_Ioi).mpr (ae_of_all _ (posf c u)), - have fpow : ∀ {c x : ℝ} (hc : 0 < c) (u : ℝ) (hx : 0 < x), - exp (-x) * x ^ (u - 1) = f c u x ^ (1 / c), - { intros c x hc u hx, - dsimp only [f], - rw [mul_rpow (exp_pos _).le ((rpow_nonneg_of_nonneg hx.le) _), ←exp_mul, ←rpow_mul hx.le], - congr' 2; - { field_simp [hc.ne'], ring } }, - -- show `f c u` is in `ℒp` for `p = 1/c`: - have f_mem_Lp : ∀ {c u : ℝ} (hc : 0 < c) (hu : 0 < u), - mem_ℒp (f c u) (ennreal.of_real (1 / c)) (volume.restrict (Ioi 0)), - { intros c u hc hu, - have A : ennreal.of_real (1 / c) ≠ 0, - by rwa [ne.def, ennreal.of_real_eq_zero, not_le, one_div_pos], - have B : ennreal.of_real (1 / c) ≠ ∞, from ennreal.of_real_ne_top, - rw [←mem_ℒp_norm_rpow_iff _ A B, ennreal.to_real_of_real (one_div_nonneg.mpr hc.le), - ennreal.div_self A B, mem_ℒp_one_iff_integrable], - { apply integrable.congr (Gamma_integral_convergent hu), - refine eventually_eq_of_mem (self_mem_ae_restrict measurable_set_Ioi) (λ x hx, _), - dsimp only, - rw fpow hc u hx, - congr' 1, - exact (norm_of_nonneg (posf _ _ x hx)).symm }, - { refine continuous_on.ae_strongly_measurable _ measurable_set_Ioi, - refine (continuous.continuous_on _).mul (continuous_at.continuous_on (λ x hx, _)), - { exact continuous_exp.comp (continuous_const.mul continuous_id'), }, - { exact continuous_at_rpow_const _ _ (or.inl (ne_of_lt hx).symm), } } }, - -- now apply Hölder: - rw [Gamma_eq_integral hs, Gamma_eq_integral ht, Gamma_eq_integral hst], - convert measure_theory.integral_mul_le_Lp_mul_Lq_of_nonneg e (posf' a s) (posf' b t) - (f_mem_Lp ha hs) (f_mem_Lp hb ht) using 1, - { refine set_integral_congr measurable_set_Ioi (λ x hx, _), - dsimp only [f], - have A : exp (-x) = exp (-a * x) * exp (-b * x), - { rw [←exp_add, ←add_mul, ←neg_add, hab, neg_one_mul] }, - have B : x ^ (a * s + b * t - 1) = (x ^ (a * (s - 1))) * (x ^ (b * (t - 1))), - { rw [←rpow_add hx, hab'], congr' 1, ring }, - rw [A, B], - ring }, - { rw [one_div_one_div, one_div_one_div], - congr' 2; - exact set_integral_congr measurable_set_Ioi (λ x hx, fpow (by assumption) _ hx) }, -end - -lemma convex_on_log_Gamma : convex_on ℝ (Ioi 0) (log ∘ Gamma) := -begin - refine convex_on_iff_forall_pos.mpr ⟨convex_Ioi _, λ x hx y hy a b ha hb hab, _⟩, - have : b = 1 - a := by linarith, subst this, - simp_rw [function.comp_app, smul_eq_mul], - rw [←log_rpow (Gamma_pos_of_pos hy), ←log_rpow (Gamma_pos_of_pos hx), - ←log_mul - ((rpow_pos_of_pos (Gamma_pos_of_pos hx) _).ne') (rpow_pos_of_pos (Gamma_pos_of_pos hy) _).ne', - log_le_log - (Gamma_pos_of_pos (add_pos (mul_pos ha hx) (mul_pos hb hy))) - (mul_pos - (rpow_pos_of_pos (Gamma_pos_of_pos hx) _) (rpow_pos_of_pos (Gamma_pos_of_pos hy) _))], - exact Gamma_mul_add_mul_le_rpow_Gamma_mul_rpow_Gamma hx hy ha hb hab, -end - -lemma convex_on_Gamma : convex_on ℝ (Ioi 0) Gamma := -begin - refine ⟨convex_Ioi 0, λ x hx y hy a b ha hb hab, _⟩, - have := convex_on.comp (convex_on_exp.subset (subset_univ _) _) convex_on_log_Gamma - (λ u hu v hv huv, exp_le_exp.mpr huv), - convert this.2 hx hy ha hb hab, - { rw [function.comp_app, exp_log (Gamma_pos_of_pos $ this.1 hx hy ha hb hab)] }, - { rw [function.comp_app, exp_log (Gamma_pos_of_pos hx)] }, - { rw [function.comp_app, exp_log (Gamma_pos_of_pos hy)] }, - { rw convex_iff_is_preconnected, - refine is_preconnected_Ioi.image _ (λ x hx, continuous_at.continuous_within_at _), - refine (differentiable_at_Gamma (λ m, _)).continuous_at.log (Gamma_pos_of_pos hx).ne', - exact (neg_lt_iff_pos_add.mpr (add_pos_of_pos_of_nonneg hx (nat.cast_nonneg m))).ne' } -end - -section bohr_mollerup - -/-! ## The Bohr-Mollerup theorem - -In this section we prove two interrelated statements about the `Γ` function on the positive reals: - -* the Euler limit formula `real.bohr_mollerup.tendsto_log_gamma_seq`, stating that for positive - real `x` the sequence `x * log n + log n! - ∑ (m : ℕ) in finset.range (n + 1), log (x + m)` - tends to `log Γ(x)` as `n → ∞`. -* the Bohr-Mollerup theorem (`real.eq_Gamma_of_log_convex`) which states that `Γ` is the unique - *log-convex*, positive-real-valued function on the positive reals satisfying - `f (x + 1) = x f x` and `f 1 = 1`. - -To do this, we prove that any function satisfying the hypotheses of the Bohr--Mollerup theorem must -agree with the limit in the Euler limit formula, so there is at most one such function. Then we -show that `Γ` satisfies these conditions. - -Since most of the auxiliary lemmas for the Bohr-Mollerup theorem are of no relevance outside the -context of this proof, we place them in a separate namespace `real.bohr_mollerup` to avoid clutter. -(This includes the logarithmic form of the Euler limit formula, since later we will prove a more -general form of the Euler limit formula valid for any real or complex `x`; see -`real.Gamma_seq_tendsto_Gamma` and `complex.Gamma_seq_tendsto_Gamma`.) --/ - -namespace bohr_mollerup - -/-- The function `n ↦ x log n + log n! - (log x + ... + log (x + n))`, which we will show tends to -`log (Gamma x)` as `n → ∞`. -/ -def log_gamma_seq (x : ℝ) (n : ℕ) : ℝ := -x * log n + log n! - ∑ (m : ℕ) in finset.range (n + 1), log (x + m) - -variables {f : ℝ → ℝ} {x : ℝ} {n : ℕ} - -lemma f_nat_eq (hf_feq : ∀ {y:ℝ}, 0 < y → f (y + 1) = f y + log y) (hn : n ≠ 0) : - f n = f 1 + log (n - 1)! := -begin - refine nat.le_induction (by simp) (λ m hm IH, _) n (nat.one_le_iff_ne_zero.2 hn), - have A : 0 < (m : ℝ), from nat.cast_pos.2 hm, - simp only [hf_feq A, nat.cast_add, algebra_map.coe_one, nat.add_succ_sub_one, add_zero], - rw [IH, add_assoc, ← log_mul (nat.cast_ne_zero.mpr (nat.factorial_ne_zero _)) A.ne', - ← nat.cast_mul], - conv_rhs { rw [← nat.succ_pred_eq_of_pos hm, nat.factorial_succ, mul_comm] }, - congr, - exact (nat.succ_pred_eq_of_pos hm).symm -end - -lemma f_add_nat_eq (hf_feq : ∀ {y:ℝ}, 0 < y → f (y + 1) = f y + log y) (hx : 0 < x) (n : ℕ) : - f (x + n) = f x + ∑ (m : ℕ) in finset.range n, log (x + m) := -begin - induction n with n hn, - { simp }, - { have : x + n.succ = (x + n) + 1, - { push_cast, ring }, - rw [this, hf_feq, hn], - rw [finset.range_succ, finset.sum_insert (finset.not_mem_range_self)], - abel, - linarith [(nat.cast_nonneg n : 0 ≤ (n:ℝ))] }, -end - -/-- Linear upper bound for `f (x + n)` on unit interval -/ -lemma f_add_nat_le - (hf_conv : convex_on ℝ (Ioi 0) f) (hf_feq : ∀ {y:ℝ}, 0 < y → f (y + 1) = f y + log y) - (hn : n ≠ 0) (hx : 0 < x) (hx' : x ≤ 1) : - f (n + x) ≤ f n + x * log n := -begin - have hn': 0 < (n:ℝ) := nat.cast_pos.mpr (nat.pos_of_ne_zero hn), - have : f n + x * log n = (1 - x) * f n + x * f (n + 1), - { rw [hf_feq hn'], ring, }, - rw [this, (by ring : (n:ℝ) + x = (1 - x) * n + x * (n + 1))], - simpa only [smul_eq_mul] using hf_conv.2 hn' (by linarith : 0 < (n + 1 : ℝ)) - (by linarith : 0 ≤ 1 - x) hx.le (by linarith), -end - -/-- Linear lower bound for `f (x + n)` on unit interval -/ -lemma f_add_nat_ge - (hf_conv : convex_on ℝ (Ioi 0) f) (hf_feq : ∀ {y:ℝ}, 0 < y → f (y + 1) = f y + log y) - (hn : 2 ≤ n) (hx : 0 < x) : - f n + x * log (n - 1) ≤ f (n + x) := -begin - have npos : 0 < (n:ℝ) - 1, - { rw [←nat.cast_one, sub_pos, nat.cast_lt], linarith, }, - have c := (convex_on_iff_slope_mono_adjacent.mp $ hf_conv).2 - npos (by linarith : 0 < (n:ℝ) + x) (by linarith : (n:ℝ) - 1 < (n:ℝ)) (by linarith), - rw [add_sub_cancel', sub_sub_cancel, div_one] at c, - have : f (↑n - 1) = f n - log (↑n - 1), - { nth_rewrite_rhs 0 (by ring : (n:ℝ) = (↑n - 1) + 1), - rw [hf_feq npos, add_sub_cancel] }, - rwa [this, le_div_iff hx, sub_sub_cancel, le_sub_iff_add_le, mul_comm _ x, add_comm] at c, -end - -lemma log_gamma_seq_add_one (x : ℝ) (n : ℕ) : - log_gamma_seq (x + 1) n = log_gamma_seq x (n + 1) + log x - (x + 1) * (log (n + 1) - log n) := -begin - dsimp only [nat.factorial_succ, log_gamma_seq], - conv_rhs { rw [finset.sum_range_succ', nat.cast_zero, add_zero], }, - rw [nat.cast_mul, log_mul], rotate, - { rw nat.cast_ne_zero, exact nat.succ_ne_zero n }, - { rw nat.cast_ne_zero, exact nat.factorial_ne_zero n, }, - have : ∑ (m : ℕ) in finset.range (n + 1), log (x + 1 + ↑m) = - ∑ (k : ℕ) in finset.range (n + 1), log (x + ↑(k + 1)), - { refine finset.sum_congr (by refl) (λ m hm, _), - congr' 1, - push_cast, - abel }, - rw [←this, nat.cast_add_one n], - ring, -end - -lemma le_log_gamma_seq - (hf_conv : convex_on ℝ (Ioi 0) f) (hf_feq : ∀ {y:ℝ}, 0 < y → f (y + 1) = f y + log y) - (hx : 0 < x) (hx' : x ≤ 1) (n : ℕ) : - f x ≤ f 1 + x * log (n + 1) - x * log n + log_gamma_seq x n := -begin - rw [log_gamma_seq, ←add_sub_assoc, le_sub_iff_add_le, ←f_add_nat_eq @hf_feq hx, add_comm x], - refine (f_add_nat_le hf_conv @hf_feq (nat.add_one_ne_zero n) hx hx').trans (le_of_eq _), - rw [f_nat_eq @hf_feq (by linarith : n + 1 ≠ 0), nat.add_sub_cancel, nat.cast_add_one], - ring, -end - -lemma ge_log_gamma_seq - (hf_conv : convex_on ℝ (Ioi 0) f) (hf_feq : ∀ {y:ℝ}, 0 < y → f (y + 1) = f y + log y) - (hx : 0 < x) (hn : n ≠ 0) : - f 1 + log_gamma_seq x n ≤ f x := -begin - dsimp [log_gamma_seq], - rw [←add_sub_assoc, sub_le_iff_le_add, ←f_add_nat_eq @hf_feq hx, add_comm x _], - refine le_trans (le_of_eq _) (f_add_nat_ge hf_conv @hf_feq _ hx), - { rw [f_nat_eq @hf_feq, nat.add_sub_cancel, nat.cast_add_one, add_sub_cancel], - { ring }, - { exact nat.succ_ne_zero _} }, - { apply nat.succ_le_succ, - linarith [nat.pos_of_ne_zero hn] }, -end - -lemma tendsto_log_gamma_seq_of_le_one - (hf_conv : convex_on ℝ (Ioi 0) f) (hf_feq : ∀ {y:ℝ}, 0 < y → f (y + 1) = f y + log y) - (hx : 0 < x) (hx' : x ≤ 1) : - tendsto (log_gamma_seq x) at_top (𝓝 $ f x - f 1) := -begin - refine tendsto_of_tendsto_of_tendsto_of_le_of_le' _ tendsto_const_nhds _ _, - show ∀ᶠ (n : ℕ) in at_top, log_gamma_seq x n ≤ f x - f 1, - { refine eventually.mp (eventually_ne_at_top 0) (eventually_of_forall (λ n hn, _)), - exact le_sub_iff_add_le'.mpr (ge_log_gamma_seq hf_conv @hf_feq hx hn) }, - show ∀ᶠ (n : ℕ) in at_top, f x - f 1 - x * (log (n + 1) - log n) ≤ log_gamma_seq x n, - { refine eventually_of_forall (λ n, _), - rw [sub_le_iff_le_add', sub_le_iff_le_add'], - convert le_log_gamma_seq hf_conv @hf_feq hx hx' n using 1, - ring }, - { have : f x - f 1 = (f x - f 1) - x * 0 := by ring, - nth_rewrite 0 this, - exact tendsto.sub tendsto_const_nhds (tendsto_log_nat_add_one_sub_log.const_mul _), } -end - -lemma tendsto_log_gamma_seq - (hf_conv : convex_on ℝ (Ioi 0) f) (hf_feq : ∀ {y:ℝ}, 0 < y → f (y + 1) = f y + log y) - (hx : 0 < x) : - tendsto (log_gamma_seq x) at_top (𝓝 $ f x - f 1) := -begin - suffices : ∀ (m : ℕ), ↑m < x → x ≤ m + 1 → - tendsto (log_gamma_seq x) at_top (𝓝 $ f x - f 1), - { refine this (⌈x - 1⌉₊) _ _, - { rcases lt_or_le x 1, - { rwa [nat.ceil_eq_zero.mpr (by linarith : x - 1 ≤ 0), nat.cast_zero] }, - { convert nat.ceil_lt_add_one (by linarith : 0 ≤ x - 1), - abel } }, - { rw ←sub_le_iff_le_add, exact nat.le_ceil _}, }, - intro m, - induction m with m hm generalizing x, - { rw [nat.cast_zero, zero_add], - exact λ _ hx', tendsto_log_gamma_seq_of_le_one hf_conv @hf_feq hx hx' }, - { intros hy hy', - rw [nat.cast_succ, ←sub_le_iff_le_add] at hy', - rw [nat.cast_succ, ←lt_sub_iff_add_lt] at hy, - specialize hm ((nat.cast_nonneg _).trans_lt hy) hy hy', - -- now massage gauss_product n (x - 1) into gauss_product (n - 1) x - have : ∀ᶠ (n:ℕ) in at_top, log_gamma_seq (x - 1) n = log_gamma_seq x (n - 1) + - x * (log (↑(n - 1) + 1) - log ↑(n - 1)) - log (x - 1), - { refine eventually.mp (eventually_ge_at_top 1) (eventually_of_forall (λ n hn, _)), - have := log_gamma_seq_add_one (x - 1) (n - 1), - rw [sub_add_cancel, nat.sub_add_cancel hn] at this, - rw this, - ring }, - replace hm := ((tendsto.congr' this hm).add - (tendsto_const_nhds : tendsto (λ _, log (x - 1)) _ _)).comp (tendsto_add_at_top_nat 1), - have : - (λ (x_1 : ℕ), (λ (n : ℕ), log_gamma_seq x (n - 1) + - x * (log (↑(n - 1) + 1) - log ↑(n - 1)) - log (x - 1)) x_1 + - (λ (b : ℕ), log (x - 1)) x_1) ∘ (λ (a : ℕ), a + 1) = - λ n, log_gamma_seq x n + x * (log (↑n + 1) - log ↑n), - { ext1 n, - dsimp only [function.comp_app], - rw [sub_add_cancel, nat.add_sub_cancel] }, - rw this at hm, - convert hm.sub (tendsto_log_nat_add_one_sub_log.const_mul x) using 2, - { ext1 n, ring }, - { have := hf_feq ((nat.cast_nonneg m).trans_lt hy), - rw sub_add_cancel at this, - rw this, - ring } }, -end - -lemma tendsto_log_Gamma {x : ℝ} (hx : 0 < x) : - tendsto (log_gamma_seq x) at_top (𝓝 $ log (Gamma x)) := -begin - have : log (Gamma x) = (log ∘ Gamma) x - (log ∘ Gamma) 1, - { simp_rw [function.comp_app, Gamma_one, log_one, sub_zero] }, - rw this, - refine bohr_mollerup.tendsto_log_gamma_seq convex_on_log_Gamma (λ y hy, _) hx, - rw [function.comp_app, Gamma_add_one hy.ne', log_mul hy.ne' (Gamma_pos_of_pos hy).ne', add_comm], -end - -end bohr_mollerup -- (namespace) - -/-- The **Bohr-Mollerup theorem**: the Gamma function is the *unique* log-convex, positive-valued -function on the positive reals which satisfies `f 1 = 1` and `f (x + 1) = x * f x` for all `x`. -/ -lemma eq_Gamma_of_log_convex {f : ℝ → ℝ} - (hf_conv : convex_on ℝ (Ioi 0) (log ∘ f)) - (hf_feq : ∀ {y:ℝ}, 0 < y → f (y + 1) = y * f y) - (hf_pos : ∀ {y:ℝ}, 0 < y → 0 < f y) - (hf_one : f 1 = 1) : - eq_on f Gamma (Ioi (0:ℝ)) := -begin - suffices : eq_on (log ∘ f) (log ∘ Gamma) (Ioi (0:ℝ)), - from λ x hx, log_inj_on_pos (hf_pos hx) (Gamma_pos_of_pos hx) (this hx), - intros x hx, - have e1 := bohr_mollerup.tendsto_log_gamma_seq hf_conv _ hx, - { rw [function.comp_app log f 1, hf_one, log_one, sub_zero] at e1, - exact tendsto_nhds_unique e1 (bohr_mollerup.tendsto_log_Gamma hx) }, - { intros y hy, - rw [function.comp_app, hf_feq hy, log_mul hy.ne' (hf_pos hy).ne'], - ring } -end - -end bohr_mollerup -- (section) - -section strict_mono - -lemma Gamma_two : Gamma 2 = 1 := by simpa using Gamma_nat_eq_factorial 1 - -lemma Gamma_three_div_two_lt_one : Gamma (3 / 2) < 1 := -begin - -- This can also be proved using the closed-form evaluation of `Gamma (1 / 2)` in - -- `analysis.special_functions.gaussian`, but we give a self-contained proof using log-convexity - -- to avoid unnecessary imports. - have A : (0:ℝ) < 3/2, by norm_num, - have := bohr_mollerup.f_add_nat_le convex_on_log_Gamma (λ y hy, _) two_ne_zero one_half_pos - (by norm_num : 1/2 ≤ (1:ℝ)), - swap, { rw [function.comp_app, Gamma_add_one hy.ne', log_mul hy.ne' (Gamma_pos_of_pos hy).ne', - add_comm] }, - rw [function.comp_app, function.comp_app, nat.cast_two, Gamma_two, log_one, zero_add, - (by norm_num : (2:ℝ) + 1/2 = 3/2 + 1), Gamma_add_one A.ne', - log_mul A.ne' (Gamma_pos_of_pos A).ne', ←le_sub_iff_add_le', - log_le_iff_le_exp (Gamma_pos_of_pos A)] at this, - refine this.trans_lt (exp_lt_one_iff.mpr _), - rw [mul_comm, ←mul_div_assoc, div_sub' _ _ (2:ℝ) two_ne_zero], - refine div_neg_of_neg_of_pos _ two_pos, - rw [sub_neg, mul_one, ←nat.cast_two, ←log_pow, ←exp_lt_exp, nat.cast_two, exp_log two_pos, - exp_log]; - norm_num, -end - -lemma Gamma_strict_mono_on_Ici : strict_mono_on Gamma (Ici 2) := -begin - convert convex_on_Gamma.strict_mono_of_lt (by norm_num : (0:ℝ) < 3/2) - (by norm_num : (3/2 : ℝ) < 2) (Gamma_two.symm ▸ Gamma_three_div_two_lt_one), - symmetry, - rw inter_eq_right_iff_subset, - exact λ x hx, two_pos.trans_le hx, -end - -end strict_mono - -end real - -section beta_integral - -/-! ## The Beta function -/ - -namespace complex - -notation `cexp` := complex.exp - -/-- The Beta function `Β (u, v)`, defined as `∫ x:ℝ in 0..1, x ^ (u - 1) * (1 - x) ^ (v - 1)`. -/ -noncomputable def beta_integral (u v : ℂ) : ℂ := -∫ (x:ℝ) in 0..1, x ^ (u - 1) * (1 - x) ^ (v - 1) - -/-- Auxiliary lemma for `beta_integral_convergent`, showing convergence at the left endpoint. -/ -lemma beta_integral_convergent_left {u : ℂ} (hu : 0 < re u) (v : ℂ) : - interval_integrable (λ x, x ^ (u - 1) * (1 - x) ^ (v - 1) : ℝ → ℂ) volume 0 (1 / 2) := -begin - apply interval_integrable.mul_continuous_on, - { refine interval_integral.interval_integrable_cpow' _, - rwa [sub_re, one_re, ←zero_sub, sub_lt_sub_iff_right] }, - { apply continuous_at.continuous_on, - intros x hx, - rw uIcc_of_le (by positivity: (0:ℝ) ≤ 1/2) at hx, - apply continuous_at.cpow, - { exact (continuous_const.sub continuous_of_real).continuous_at }, - { exact continuous_at_const }, - { rw [sub_re, one_re, of_real_re, sub_pos], - exact or.inl (hx.2.trans_lt (by norm_num : (1/2:ℝ) < 1)) } } -end - -/-- The Beta integral is convergent for all `u, v` of positive real part. -/ -lemma beta_integral_convergent {u v : ℂ} (hu : 0 < re u) (hv : 0 < re v) : - interval_integrable (λ x, x ^ (u - 1) * (1 - x) ^ (v - 1) : ℝ → ℂ) volume 0 1 := -begin - refine (beta_integral_convergent_left hu v).trans _, - rw interval_integrable.iff_comp_neg, - convert ((beta_integral_convergent_left hv u).comp_add_right 1).symm, - { ext1 x, - conv_lhs { rw mul_comm }, - congr' 2; - { push_cast, ring } }, - { norm_num }, - { norm_num } -end - -lemma beta_integral_symm (u v : ℂ) : - beta_integral v u = beta_integral u v := -begin - rw [beta_integral, beta_integral], - have := interval_integral.integral_comp_mul_add - (λ x:ℝ, (x:ℂ) ^ (u - 1) * (1 - ↑x) ^ (v - 1)) (neg_one_lt_zero.ne) 1, - rw [inv_neg, inv_one, neg_one_smul, ←interval_integral.integral_symm] at this, - convert this, - { ext1 x, rw mul_comm, congr; - { push_cast, ring } }, - { ring }, { ring } -end - -lemma beta_integral_eval_one_right {u : ℂ} (hu : 0 < re u) : - beta_integral u 1 = 1 / u := -begin - simp_rw [beta_integral, sub_self, cpow_zero, mul_one], - rw integral_cpow (or.inl _), - { rw [of_real_zero, of_real_one, one_cpow, zero_cpow, - sub_zero, sub_add_cancel], - rw sub_add_cancel, - contrapose! hu, rw [hu, zero_re] }, - { rwa [sub_re, one_re, ←sub_pos, sub_neg_eq_add, sub_add_cancel] }, -end - -lemma beta_integral_scaled (s t : ℂ) {a : ℝ} (ha : 0 < a) : - ∫ x in 0..a, (x:ℂ) ^ (s - 1) * (a - x) ^ (t - 1) = a ^ (s + t - 1) * beta_integral s t := -begin - have ha' : (a:ℂ) ≠ 0, from of_real_ne_zero.mpr ha.ne', - rw beta_integral, - have A : (a:ℂ) ^ (s + t - 1) = a * (a ^ (s - 1) * a ^ (t - 1)), - { rw [(by abel : s + t - 1 = 1 + (s - 1) + (t - 1)), - cpow_add _ _ ha', cpow_add 1 _ ha', cpow_one, mul_assoc] }, - rw [A, mul_assoc, ←interval_integral.integral_const_mul ((↑a) ^ _ * _), - ←real_smul, ←(zero_div a), ←div_self ha.ne', - ←interval_integral.integral_comp_div _ ha.ne', zero_div], - simp_rw interval_integral.integral_of_le ha.le, - refine set_integral_congr measurable_set_Ioc (λ x hx, _), - dsimp only, - rw mul_mul_mul_comm, - congr' 1, - { rw [←mul_cpow_of_real_nonneg ha.le (div_pos hx.1 ha).le, of_real_div, mul_div_cancel' _ ha'] }, - { rw [(by push_cast : (1:ℂ) - ↑(x / a) = ↑(1 - x / a)), - ←mul_cpow_of_real_nonneg ha.le (sub_nonneg.mpr $ (div_le_one ha).mpr hx.2)], - push_cast, - rw [mul_sub, mul_one, mul_div_cancel' _ ha'] } -end - -/-- Relation between Beta integral and Gamma function. -/ -lemma Gamma_mul_Gamma_eq_beta_integral {s t : ℂ} (hs : 0 < re s) (ht : 0 < re t) : - Gamma s * Gamma t = Gamma (s + t) * beta_integral s t := -begin - -- Note that we haven't proved (yet) that the Gamma function has no zeroes, so we can't formulate - -- this as a formula for the Beta function. - have conv_int := integral_pos_convolution (Gamma_integral_convergent hs) - (Gamma_integral_convergent ht) (continuous_linear_map.mul ℝ ℂ), - simp_rw continuous_linear_map.mul_apply' at conv_int, - have hst : 0 < re (s + t), - { rw add_re, exact add_pos hs ht }, - rw [Gamma_eq_integral hs, Gamma_eq_integral ht, Gamma_eq_integral hst, Gamma_integral, - Gamma_integral, Gamma_integral, ←conv_int, ←integral_mul_right (beta_integral _ _)], - refine set_integral_congr measurable_set_Ioi (λ x hx, _), - dsimp only, - rw [mul_assoc, ←beta_integral_scaled s t hx, ←interval_integral.integral_const_mul], - congr' 1 with y:1, - push_cast, - suffices : cexp (-x) = cexp (-y) * cexp (-(x - y)), - { rw this, ring }, - { rw ←complex.exp_add, congr' 1, abel }, -end - -/-- Recurrence formula for the Beta function. -/ -lemma beta_integral_recurrence {u v : ℂ} (hu : 0 < re u) (hv : 0 < re v) : - u * beta_integral u (v + 1) = v * beta_integral (u + 1) v := -begin - -- NB: If we knew `Gamma (u + v + 1) ≠ 0` this would be an easy consequence of - -- `Gamma_mul_Gamma_eq_beta_integral`; but we don't know that yet. We will prove it later, but - -- this lemma is needed in the proof. So we give a (somewhat laborious) direct argument. - let F : ℝ → ℂ := λ x, x ^ u * (1 - x) ^ v, - have hu' : 0 < re (u + 1), by { rw [add_re, one_re], positivity }, - have hv' : 0 < re (v + 1), by { rw [add_re, one_re], positivity }, - have hc : continuous_on F (Icc 0 1), - { refine (continuous_at.continuous_on (λ x hx, _)).mul (continuous_at.continuous_on (λ x hx, _)), - { refine (continuous_at_cpow_const_of_re_pos (or.inl _) hu).comp - continuous_of_real.continuous_at, - rw of_real_re, exact hx.1 }, - { refine (continuous_at_cpow_const_of_re_pos (or.inl _) hv).comp - (continuous_const.sub continuous_of_real).continuous_at, - rw [sub_re, one_re, of_real_re, sub_nonneg], - exact hx.2 } }, - have hder : ∀ (x : ℝ), x ∈ Ioo (0:ℝ) 1 → has_deriv_at F - (u * (↑x ^ (u - 1) * (1 - ↑x) ^ v) - v * (↑x ^ u * (1 - ↑x) ^ (v - 1))) x, - { intros x hx, - have U : has_deriv_at (λ y:ℂ, y ^ u) (u * ↑x ^ (u - 1)) ↑x, - { have := has_deriv_at.cpow_const (has_deriv_at_id ↑x) (or.inl _), - { rw mul_one at this, exact this }, - { rw [id.def, of_real_re], exact hx.1 } }, - have V : has_deriv_at (λ y:ℂ, (1 - y) ^ v) (-v * (1 - ↑x) ^ (v - 1)) ↑x, - { have A := has_deriv_at.cpow_const (has_deriv_at_id (1 - ↑x)) (or.inl _), - rotate, { exact v }, - { rw [id.def, sub_re, one_re, of_real_re, sub_pos], exact hx.2 }, - simp_rw [id.def] at A, - have B : has_deriv_at (λ y:ℂ, 1 - y) (-1) ↑x, - { apply has_deriv_at.const_sub, apply has_deriv_at_id }, - convert has_deriv_at.comp ↑x A B using 1, - ring }, - convert (U.mul V).comp_of_real, - ring }, - have h_int := ((beta_integral_convergent hu hv').const_mul u).sub - ((beta_integral_convergent hu' hv).const_mul v), - dsimp only at h_int, - rw [add_sub_cancel, add_sub_cancel] at h_int, - have int_ev := interval_integral.integral_eq_sub_of_has_deriv_at_of_le zero_le_one hc hder h_int, - have hF0 : F 0 = 0, - { simp only [mul_eq_zero, of_real_zero, cpow_eq_zero_iff, eq_self_iff_true, - ne.def, true_and, sub_zero, one_cpow, one_ne_zero, or_false], - contrapose! hu, rw [hu, zero_re] }, - have hF1 : F 1 = 0, - { simp only [mul_eq_zero, of_real_one, one_cpow, one_ne_zero, sub_self, - cpow_eq_zero_iff, eq_self_iff_true, ne.def, true_and, false_or], - contrapose! hv, rw [hv, zero_re] }, - rw [hF0, hF1, sub_zero, interval_integral.integral_sub, - interval_integral.integral_const_mul, interval_integral.integral_const_mul] at int_ev, - { rw [beta_integral, beta_integral, ←sub_eq_zero], - convert int_ev; - { ext1 x, congr, abel } }, - { apply interval_integrable.const_mul, - convert beta_integral_convergent hu hv', - ext1 x, rw add_sub_cancel }, - { apply interval_integrable.const_mul, - convert beta_integral_convergent hu' hv, - ext1 x, rw add_sub_cancel }, -end - -/-- Explicit formula for the Beta function when second argument is a positive integer. -/ -lemma beta_integral_eval_nat_add_one_right {u : ℂ} (hu : 0 < re u) (n : ℕ) : - beta_integral u (n + 1) = n! / ∏ (j:ℕ) in finset.range (n + 1), (u + j) := -begin - induction n with n IH generalizing u, - { rw [nat.cast_zero, zero_add, beta_integral_eval_one_right hu, - nat.factorial_zero, nat.cast_one, zero_add, finset.prod_range_one, nat.cast_zero, add_zero] }, - { have := beta_integral_recurrence hu (_ : 0 < re n.succ), - swap, { rw [←of_real_nat_cast, of_real_re], positivity }, - rw [mul_comm u _, ←eq_div_iff] at this, - swap, { contrapose! hu, rw [hu, zero_re] }, - rw [this, finset.prod_range_succ', nat.cast_succ, IH], - swap, { rw [add_re, one_re], positivity }, - rw [nat.factorial_succ, nat.cast_mul, nat.cast_add, nat.cast_one, nat.cast_zero, add_zero, - ←mul_div_assoc, ←div_div], - congr' 3 with j:1, - push_cast, abel } -end - -end complex - -end beta_integral - -section limit_formula - -/-! ## The Euler limit formula -/ - -namespace complex - -/-- The sequence with `n`-th term `n ^ s * n! / (s * (s + 1) * ... * (s + n))`, for complex `s`. -We will show that this tends to `Γ(s)` as `n → ∞`. -/ -noncomputable def Gamma_seq (s : ℂ) (n : ℕ) := -(n:ℂ) ^ s * n! / ∏ (j:ℕ) in finset.range (n + 1), (s + j) - -lemma Gamma_seq_eq_beta_integral_of_re_pos {s : ℂ} (hs : 0 < re s) (n : ℕ) : - Gamma_seq s n = n ^ s * beta_integral s (n + 1) := -by rw [Gamma_seq, beta_integral_eval_nat_add_one_right hs n, ←mul_div_assoc] - -lemma Gamma_seq_add_one_left (s : ℂ) {n : ℕ} (hn : n ≠ 0) : - (Gamma_seq (s + 1) n) / s = n / (n + 1 + s) * Gamma_seq s n := -begin - conv_lhs { rw [Gamma_seq, finset.prod_range_succ, div_div] }, - conv_rhs { rw [Gamma_seq, finset.prod_range_succ', nat.cast_zero, add_zero, div_mul_div_comm, - ←mul_assoc, ←mul_assoc, mul_comm _ (finset.prod _ _)] }, - congr' 3, - { rw [cpow_add _ _ (nat.cast_ne_zero.mpr hn), cpow_one, mul_comm] }, - { refine finset.prod_congr (by refl) (λ x hx, _), - push_cast, ring }, - { abel } -end - -lemma Gamma_seq_eq_approx_Gamma_integral {s : ℂ} (hs : 0 < re s) {n : ℕ} (hn : n ≠ 0) : - Gamma_seq s n = ∫ x:ℝ in 0..n, ↑((1 - x / n) ^ n) * (x:ℂ) ^ (s - 1) := -begin - have : ∀ (x : ℝ), x = x / n * n, by { intro x, rw div_mul_cancel, exact nat.cast_ne_zero.mpr hn }, - conv in (↑_ ^ _) { congr, rw this x }, - rw Gamma_seq_eq_beta_integral_of_re_pos hs, - rw [beta_integral, @interval_integral.integral_comp_div _ _ _ _ 0 n _ - (λ x, ↑((1 - x) ^ n) * ↑(x * ↑n) ^ (s - 1) : ℝ → ℂ) (nat.cast_ne_zero.mpr hn), - real_smul, zero_div, div_self, add_sub_cancel, ←interval_integral.integral_const_mul, - ←interval_integral.integral_const_mul], - swap, { exact nat.cast_ne_zero.mpr hn }, - simp_rw interval_integral.integral_of_le zero_le_one, - refine set_integral_congr measurable_set_Ioc (λ x hx, _), - push_cast, - have hn' : (n : ℂ) ≠ 0, from nat.cast_ne_zero.mpr hn, - have A : (n : ℂ) ^ s = (n : ℂ) ^ (s - 1) * n, - { conv_lhs { rw [(by ring : s = (s - 1) + 1), cpow_add _ _ hn'] }, - simp }, - have B : ((x : ℂ) * ↑n) ^ (s - 1) = (x : ℂ) ^ (s - 1) * ↑n ^ (s - 1), - { rw [←of_real_nat_cast, - mul_cpow_of_real_nonneg hx.1.le (nat.cast_pos.mpr (nat.pos_of_ne_zero hn)).le] }, - rw [A, B, cpow_nat_cast], ring, -end - -/-- The main techical lemma for `Gamma_seq_tendsto_Gamma`, expressing the integral defining the -Gamma function for `0 < re s` as the limit of a sequence of integrals over finite intervals. -/ -lemma approx_Gamma_integral_tendsto_Gamma_integral {s : ℂ} (hs : 0 < re s) : - tendsto (λ n:ℕ, ∫ x:ℝ in 0..n, ↑((1 - x / n) ^ n) * (x:ℂ) ^ (s - 1)) at_top (𝓝 $ Gamma s) := -begin - rw [Gamma_eq_integral hs], - -- We apply dominated convergence to the following function, which we will show is uniformly - -- bounded above by the Gamma integrand `exp (-x) * x ^ (re s - 1)`. - let f : ℕ → ℝ → ℂ := λ n, indicator (Ioc 0 (n:ℝ)) - (λ x:ℝ, ↑((1 - x / n) ^ n) * (x:ℂ) ^ (s - 1)), - -- integrability of f - have f_ible : ∀ (n:ℕ), integrable (f n) (volume.restrict (Ioi 0)), - { intro n, - rw [integrable_indicator_iff (measurable_set_Ioc : measurable_set (Ioc (_:ℝ) _)), - integrable_on, measure.restrict_restrict_of_subset Ioc_subset_Ioi_self, ←integrable_on, - ←interval_integrable_iff_integrable_Ioc_of_le (by positivity : (0:ℝ) ≤ n)], - apply interval_integrable.continuous_on_mul, - { refine interval_integral.interval_integrable_cpow' _, - rwa [sub_re, one_re, ←zero_sub, sub_lt_sub_iff_right] }, - { apply continuous.continuous_on, continuity } }, - -- pointwise limit of f - have f_tends : ∀ x:ℝ, x ∈ Ioi (0:ℝ) → - tendsto (λ n:ℕ, f n x) at_top (𝓝 $ ↑(real.exp (-x)) * (x:ℂ) ^ (s - 1)), - { intros x hx, - apply tendsto.congr', - show ∀ᶠ n:ℕ in at_top, ↑((1 - x / n) ^ n) * (x:ℂ) ^ (s - 1) = f n x, - { refine eventually.mp (eventually_ge_at_top ⌈x⌉₊) (eventually_of_forall (λ n hn, _)), - rw nat.ceil_le at hn, - dsimp only [f], - rw indicator_of_mem, - exact ⟨hx, hn⟩ }, - { simp_rw mul_comm _ (↑x ^ _), - refine (tendsto.comp (continuous_of_real.tendsto _) _).const_mul _, - convert tendsto_one_plus_div_pow_exp (-x), - ext1 n, - rw [neg_div, ←sub_eq_add_neg] } }, - -- let `convert` identify the remaining goals - convert tendsto_integral_of_dominated_convergence _ (λ n, (f_ible n).1) - (real.Gamma_integral_convergent hs) _ - ((ae_restrict_iff' measurable_set_Ioi).mpr (ae_of_all _ f_tends)), - -- limit of f is the integrand we want - { ext1 n, - rw [integral_indicator (measurable_set_Ioc : measurable_set (Ioc (_:ℝ) _)), - interval_integral.integral_of_le (by positivity: 0 ≤ (n:ℝ)), - measure.restrict_restrict_of_subset Ioc_subset_Ioi_self] }, - -- f is uniformly bounded by the Gamma integrand - { intro n, - refine (ae_restrict_iff' measurable_set_Ioi).mpr (ae_of_all _ (λ x hx, _)), - dsimp only [f], - rcases lt_or_le (n:ℝ) x with hxn | hxn, - { rw [indicator_of_not_mem (not_mem_Ioc_of_gt hxn), norm_zero, - mul_nonneg_iff_right_nonneg_of_pos (exp_pos _)], - exact rpow_nonneg_of_nonneg (le_of_lt hx) _ }, - { rw [indicator_of_mem (mem_Ioc.mpr ⟨hx, hxn⟩), norm_mul, complex.norm_eq_abs, - complex.abs_of_nonneg - (pow_nonneg (sub_nonneg.mpr $ div_le_one_of_le hxn $ by positivity) _), - complex.norm_eq_abs, abs_cpow_eq_rpow_re_of_pos hx, sub_re, one_re, - mul_le_mul_right (rpow_pos_of_pos hx _ )], - exact one_sub_div_pow_le_exp_neg hxn } } -end - -/-- Euler's limit formula for the complex Gamma function. -/ -lemma Gamma_seq_tendsto_Gamma (s : ℂ) : - tendsto (Gamma_seq s) at_top (𝓝 $ Gamma s) := -begin - suffices : ∀ m : ℕ, (-↑m < re s) → tendsto (Gamma_seq s) at_top (𝓝 $ Gamma_aux m s), - { rw Gamma, - apply this, - rw neg_lt, - rcases lt_or_le 0 (re s) with hs | hs, - { exact (neg_neg_of_pos hs).trans_le (nat.cast_nonneg _), }, - { refine (nat.lt_floor_add_one _).trans_le _, - rw [sub_eq_neg_add, nat.floor_add_one (neg_nonneg.mpr hs), nat.cast_add_one] } }, - intro m, - induction m with m IH generalizing s, - { -- Base case: `0 < re s`, so Gamma is given by the integral formula - intro hs, - rw [nat.cast_zero, neg_zero] at hs, - rw [←Gamma_eq_Gamma_aux], - { refine tendsto.congr' _ (approx_Gamma_integral_tendsto_Gamma_integral hs), - refine (eventually_ne_at_top 0).mp (eventually_of_forall (λ n hn, _)), - exact (Gamma_seq_eq_approx_Gamma_integral hs hn).symm }, - { rwa [nat.cast_zero, neg_lt_zero] } }, - { -- Induction step: use recurrence formulae in `s` for Gamma and Gamma_seq - intro hs, - rw [nat.cast_succ, neg_add, ←sub_eq_add_neg, sub_lt_iff_lt_add, ←one_re, ←add_re] at hs, - rw Gamma_aux, - have := tendsto.congr' ((eventually_ne_at_top 0).mp (eventually_of_forall (λ n hn, _))) - ((IH _ hs).div_const s), - swap 3, { exact Gamma_seq_add_one_left s hn }, -- doesn't work if inlined? - conv at this in (_ / _ * _) { rw mul_comm }, - rwa [←mul_one (Gamma_aux m (s + 1) / s), tendsto_mul_iff_of_ne_zero _ (one_ne_zero' ℂ)] at this, - simp_rw add_assoc, - exact tendsto_coe_nat_div_add_at_top (1 + s) } -end - -end complex - -end limit_formula - -section gamma_reflection -/-! ## The reflection formula -/ - -open_locale real -namespace complex - -lemma Gamma_seq_mul (z : ℂ) {n : ℕ} (hn : n ≠ 0) : - Gamma_seq z n * Gamma_seq (1 - z) n = - n / (n + 1 - z) * (1 / (z * ∏ j in finset.range n, (1 - z ^ 2 / (j + 1) ^ 2))) := -begin - -- also true for n = 0 but we don't need it - have aux : ∀ (a b c d : ℂ), a * b * (c * d) = a * c * (b * d), by { intros, ring }, - rw [Gamma_seq, Gamma_seq, div_mul_div_comm, aux, ←pow_two], - have : (n : ℂ) ^ z * n ^ (1 - z) = n, - { rw [←cpow_add _ _ (nat.cast_ne_zero.mpr hn), add_sub_cancel'_right, cpow_one] }, - rw [this, finset.prod_range_succ', finset.prod_range_succ, aux, ←finset.prod_mul_distrib, - nat.cast_zero, add_zero, add_comm (1 - z) n, ←add_sub_assoc], - have : ∀ (j : ℕ), (z + ↑(j + 1)) * (1 - z + ↑j) = ↑((j + 1) ^ 2) * (1 - z ^ 2 / (↑j + 1) ^ 2), - { intro j, - push_cast, - have : (j:ℂ) + 1 ≠ 0, by { rw [←nat.cast_succ, nat.cast_ne_zero], exact nat.succ_ne_zero j }, - field_simp, ring }, - simp_rw this, - rw [finset.prod_mul_distrib, ←nat.cast_prod, finset.prod_pow, - finset.prod_range_add_one_eq_factorial, nat.cast_pow, - (by {intros, ring} : ∀ (a b c d : ℂ), a * b * (c * d) = a * (d * (b * c))), - ←div_div, mul_div_cancel, ←div_div, mul_comm z _, mul_one_div], - exact pow_ne_zero 2 (nat.cast_ne_zero.mpr $ nat.factorial_ne_zero n), -end - -/-- Euler's reflection formula for the complex Gamma function. -/ -theorem Gamma_mul_Gamma_one_sub (z : ℂ) : Gamma z * Gamma (1 - z) = π / sin (π * z) := -begin - have pi_ne : (π : ℂ) ≠ 0, from complex.of_real_ne_zero.mpr pi_ne_zero, - by_cases hs : sin (↑π * z) = 0, - { -- first deal with silly case z = integer - rw [hs, div_zero], - rw [←neg_eq_zero, ←complex.sin_neg, ←mul_neg, complex.sin_eq_zero_iff, mul_comm] at hs, - obtain ⟨k, hk⟩ := hs, - rw [mul_eq_mul_right_iff, eq_false_intro (of_real_ne_zero.mpr pi_pos.ne'), or_false, - neg_eq_iff_eq_neg] at hk, - rw hk, - cases k, - { rw [int.cast_of_nat, complex.Gamma_neg_nat_eq_zero, zero_mul] }, - { rw [int.cast_neg_succ_of_nat, neg_neg, nat.cast_add, nat.cast_one, add_comm, sub_add_cancel', - complex.Gamma_neg_nat_eq_zero, mul_zero] } }, - refine tendsto_nhds_unique ((Gamma_seq_tendsto_Gamma z).mul (Gamma_seq_tendsto_Gamma $ 1 - z)) _, - have : ↑π / sin (↑π * z) = 1 * (π / sin (π * z)), by rw one_mul, rw this, - refine tendsto.congr' ((eventually_ne_at_top 0).mp - (eventually_of_forall (λ n hn, (Gamma_seq_mul z hn).symm))) (tendsto.mul _ _), - { convert tendsto_coe_nat_div_add_at_top (1 - z), ext1 n, rw add_sub_assoc }, - { have : ↑π / sin (↑π * z) = 1 / (sin (π * z) / π), by field_simp, rw this, - refine tendsto_const_nhds.div _ (div_ne_zero hs pi_ne), - rw [←tendsto_mul_iff_of_ne_zero tendsto_const_nhds pi_ne, div_mul_cancel _ pi_ne], - convert tendsto_euler_sin_prod z, - ext1 n, rw [mul_comm, ←mul_assoc] }, -end - -/-- The Gamma function does not vanish on `ℂ` (except at non-positive integers, where the function -is mathematically undefined and we set it to `0` by convention). -/ -theorem Gamma_ne_zero {s : ℂ} (hs : ∀ m : ℕ, s ≠ -m) : Gamma s ≠ 0 := -begin - by_cases h_im : s.im = 0, - { have : s = ↑s.re, - { conv_lhs { rw ←complex.re_add_im s }, rw [h_im, of_real_zero, zero_mul, add_zero] }, - rw [this, Gamma_of_real, of_real_ne_zero], - refine real.Gamma_ne_zero (λ n, _), - specialize hs n, - contrapose! hs, - rwa [this, ←of_real_nat_cast, ←of_real_neg, of_real_inj] }, - { have : sin (↑π * s) ≠ 0, - { rw complex.sin_ne_zero_iff, - intro k, - apply_fun im, - rw [of_real_mul_im, ←of_real_int_cast, ←of_real_mul, of_real_im], - exact mul_ne_zero real.pi_pos.ne' h_im }, - have A := div_ne_zero (of_real_ne_zero.mpr real.pi_pos.ne') this, - rw [←complex.Gamma_mul_Gamma_one_sub s, mul_ne_zero_iff] at A, - exact A.1 } -end - -lemma Gamma_eq_zero_iff (s : ℂ) : Gamma s = 0 ↔ ∃ m : ℕ, s = -m := -begin - split, - { contrapose!, exact Gamma_ne_zero }, - { rintro ⟨m, rfl⟩, exact Gamma_neg_nat_eq_zero m }, -end - -end complex - -namespace real - -/-- The sequence with `n`-th term `n ^ s * n! / (s * (s + 1) * ... * (s + n))`, for real `s`. We -will show that this tends to `Γ(s)` as `n → ∞`. -/ -noncomputable def Gamma_seq (s : ℝ) (n : ℕ) := -(n : ℝ) ^ s * n! / ∏ (j : ℕ) in finset.range (n + 1), (s + j) - -/-- Euler's limit formula for the real Gamma function. -/ -lemma Gamma_seq_tendsto_Gamma (s : ℝ) : tendsto (Gamma_seq s) at_top (𝓝 $ Gamma s) := -begin - suffices : tendsto (coe ∘ Gamma_seq s : ℕ → ℂ) at_top (𝓝 $ complex.Gamma s), - from (complex.continuous_re.tendsto (complex.Gamma ↑s)).comp this, - convert complex.Gamma_seq_tendsto_Gamma s, - ext1 n, - dsimp only [Gamma_seq, function.comp_app, complex.Gamma_seq], - push_cast, - rw [complex.of_real_cpow n.cast_nonneg, complex.of_real_nat_cast] -end - -/-- Euler's reflection formula for the real Gamma function. -/ -lemma Gamma_mul_Gamma_one_sub (s : ℝ) : Gamma s * Gamma (1 - s) = π / sin (π * s) := -begin - simp_rw [←complex.of_real_inj, complex.of_real_div, complex.of_real_sin, - complex.of_real_mul, ←complex.Gamma_of_real, complex.of_real_sub, complex.of_real_one], - exact complex.Gamma_mul_Gamma_one_sub s -end - -end real - -end gamma_reflection diff --git a/src/analysis/special_functions/gamma/basic.lean b/src/analysis/special_functions/gamma/basic.lean new file mode 100644 index 0000000000000..1d488c767c51c --- /dev/null +++ b/src/analysis/special_functions/gamma/basic.lean @@ -0,0 +1,673 @@ +/- +Copyright (c) 2022 David Loeffler. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: David Loeffler +-/ +import measure_theory.integral.exp_decay +import analysis.special_functions.improper_integrals +import analysis.calculus.parametric_integral + +/-! +# The Gamma function + +This file defines the `Γ` function (of a real or complex variable `s`). We define this by Euler's +integral `Γ(s) = ∫ x in Ioi 0, exp (-x) * x ^ (s - 1)` in the range where this integral converges +(i.e., for `0 < s` in the real case, and `0 < re s` in the complex case). + +We show that this integral satisfies `Γ(1) = 1` and `Γ(s + 1) = s * Γ(s)`; hence we can define +`Γ(s)` for all `s` as the unique function satisfying this recurrence and agreeing with Euler's +integral in the convergence range. (If `s = -n` for `n ∈ ℕ`, then the function is undefined, and we +set it to be `0` by convention.) + +## Gamma function: main statements (complex case) + +* `complex.Gamma`: the `Γ` function (of a complex variable). +* `complex.Gamma_eq_integral`: for `0 < re s`, `Γ(s)` agrees with Euler's integral. +* `complex.Gamma_add_one`: for all `s : ℂ` with `s ≠ 0`, we have `Γ (s + 1) = s Γ(s)`. +* `complex.Gamma_nat_eq_factorial`: for all `n : ℕ` we have `Γ (n + 1) = n!`. +* `complex.differentiable_at_Gamma`: `Γ` is complex-differentiable at all `s : ℂ` with + `s ∉ {-n : n ∈ ℕ}`. + +## Gamma function: main statements (real case) + +* `real.Gamma`: the `Γ` function (of a real variable). +* Real counterparts of all the properties of the complex Gamma function listed above: + `real.Gamma_eq_integral`, `real.Gamma_add_one`, `real.Gamma_nat_eq_factorial`, + `real.differentiable_at_Gamma`. + +## Tags + +Gamma +-/ + +noncomputable theory +open filter interval_integral set real measure_theory asymptotics +open_locale nat topology complex_conjugate + +namespace real + +/-- Asymptotic bound for the `Γ` function integrand. -/ +lemma Gamma_integrand_is_o (s : ℝ) : + (λ x:ℝ, exp (-x) * x ^ s) =o[at_top] (λ x:ℝ, exp (-(1/2) * x)) := +begin + refine is_o_of_tendsto (λ x hx, _) _, + { exfalso, exact (exp_pos (-(1 / 2) * x)).ne' hx }, + have : (λ (x:ℝ), exp (-x) * x ^ s / exp (-(1 / 2) * x)) = (λ (x:ℝ), exp ((1 / 2) * x) / x ^ s )⁻¹, + { ext1 x, + field_simp [exp_ne_zero, exp_neg, ← real.exp_add], + left, + ring }, + rw this, + exact (tendsto_exp_mul_div_rpow_at_top s (1 / 2) one_half_pos).inv_tendsto_at_top, +end + +/-- The Euler integral for the `Γ` function converges for positive real `s`. -/ +lemma Gamma_integral_convergent {s : ℝ} (h : 0 < s) : + integrable_on (λ x:ℝ, exp (-x) * x ^ (s - 1)) (Ioi 0) := +begin + rw [←Ioc_union_Ioi_eq_Ioi (@zero_le_one ℝ _ _ _ _), integrable_on_union], + split, + { rw ←integrable_on_Icc_iff_integrable_on_Ioc, + refine integrable_on.continuous_on_mul continuous_on_id.neg.exp _ is_compact_Icc, + refine (interval_integrable_iff_integrable_Icc_of_le zero_le_one).mp _, + exact interval_integrable_rpow' (by linarith), }, + { refine integrable_of_is_O_exp_neg one_half_pos _ (Gamma_integrand_is_o _ ).is_O, + refine continuous_on_id.neg.exp.mul (continuous_on_id.rpow_const _), + intros x hx, + exact or.inl ((zero_lt_one : (0 : ℝ) < 1).trans_le hx).ne' } +end + +end real + +namespace complex +/- Technical note: In defining the Gamma integrand exp (-x) * x ^ (s - 1) for s complex, we have to +make a choice between ↑(real.exp (-x)), complex.exp (↑(-x)), and complex.exp (-↑x), all of which are +equal but not definitionally so. We use the first of these throughout. -/ + + +/-- The integral defining the `Γ` function converges for complex `s` with `0 < re s`. + +This is proved by reduction to the real case. -/ +lemma Gamma_integral_convergent {s : ℂ} (hs : 0 < s.re) : + integrable_on (λ x, (-x).exp * x ^ (s - 1) : ℝ → ℂ) (Ioi 0) := +begin + split, + { refine continuous_on.ae_strongly_measurable _ measurable_set_Ioi, + apply (continuous_of_real.comp continuous_neg.exp).continuous_on.mul, + apply continuous_at.continuous_on, + intros x hx, + have : continuous_at (λ x:ℂ, x ^ (s - 1)) ↑x, + { apply continuous_at_cpow_const, rw of_real_re, exact or.inl hx, }, + exact continuous_at.comp this continuous_of_real.continuous_at }, + { rw ←has_finite_integral_norm_iff, + refine has_finite_integral.congr (real.Gamma_integral_convergent hs).2 _, + refine (ae_restrict_iff' measurable_set_Ioi).mpr (ae_of_all _ (λ x hx, _)), + dsimp only, + rw [norm_eq_abs, map_mul, abs_of_nonneg $ le_of_lt $ exp_pos $ -x, + abs_cpow_eq_rpow_re_of_pos hx _], + simp } +end + +/-- Euler's integral for the `Γ` function (of a complex variable `s`), defined as +`∫ x in Ioi 0, exp (-x) * x ^ (s - 1)`. + +See `complex.Gamma_integral_convergent` for a proof of the convergence of the integral for +`0 < re s`. -/ +def Gamma_integral (s : ℂ) : ℂ := ∫ x in Ioi (0:ℝ), ↑(-x).exp * ↑x ^ (s - 1) + +lemma Gamma_integral_conj (s : ℂ) : Gamma_integral (conj s) = conj (Gamma_integral s) := +begin + rw [Gamma_integral, Gamma_integral, ←integral_conj], + refine set_integral_congr measurable_set_Ioi (λ x hx, _), + dsimp only, + rw [ring_hom.map_mul, conj_of_real, cpow_def_of_ne_zero (of_real_ne_zero.mpr (ne_of_gt hx)), + cpow_def_of_ne_zero (of_real_ne_zero.mpr (ne_of_gt hx)), ←exp_conj, ring_hom.map_mul, + ←of_real_log (le_of_lt hx), conj_of_real, ring_hom.map_sub, ring_hom.map_one], +end + +lemma Gamma_integral_of_real (s : ℝ) : + Gamma_integral ↑s = ↑(∫ x:ℝ in Ioi 0, real.exp (-x) * x ^ (s - 1)) := +begin + rw [Gamma_integral, ←_root_.integral_of_real], + refine set_integral_congr measurable_set_Ioi _, + intros x hx, dsimp only, + rw [of_real_mul, of_real_cpow (mem_Ioi.mp hx).le], + simp, +end + +lemma Gamma_integral_one : Gamma_integral 1 = 1 := +by simpa only [←of_real_one, Gamma_integral_of_real, of_real_inj, sub_self, + rpow_zero, mul_one] using integral_exp_neg_Ioi_zero + +end complex + +/-! Now we establish the recurrence relation `Γ(s + 1) = s * Γ(s)` using integration by parts. -/ + +namespace complex + +section Gamma_recurrence + +/-- The indefinite version of the `Γ` function, `Γ(s, X) = ∫ x ∈ 0..X, exp(-x) x ^ (s - 1)`. -/ +def partial_Gamma (s : ℂ) (X : ℝ) : ℂ := ∫ x in 0..X, (-x).exp * x ^ (s - 1) + +lemma tendsto_partial_Gamma {s : ℂ} (hs: 0 < s.re) : + tendsto (λ X:ℝ, partial_Gamma s X) at_top (𝓝 $ Gamma_integral s) := +interval_integral_tendsto_integral_Ioi 0 (Gamma_integral_convergent hs) tendsto_id + +private lemma Gamma_integrand_interval_integrable (s : ℂ) {X : ℝ} (hs : 0 < s.re) (hX : 0 ≤ X): + interval_integrable (λ x, (-x).exp * x ^ (s - 1) : ℝ → ℂ) volume 0 X := +begin + rw interval_integrable_iff_integrable_Ioc_of_le hX, + exact integrable_on.mono_set (Gamma_integral_convergent hs) Ioc_subset_Ioi_self +end + +private lemma Gamma_integrand_deriv_integrable_A {s : ℂ} (hs : 0 < s.re) {X : ℝ} (hX : 0 ≤ X): + interval_integrable (λ x, -((-x).exp * x ^ s) : ℝ → ℂ) volume 0 X := +begin + convert (Gamma_integrand_interval_integrable (s+1) _ hX).neg, + { ext1, simp only [add_sub_cancel, pi.neg_apply] }, + { simp only [add_re, one_re], linarith,}, +end + +private lemma Gamma_integrand_deriv_integrable_B {s : ℂ} (hs : 0 < s.re) {Y : ℝ} (hY : 0 ≤ Y) : + interval_integrable (λ (x : ℝ), (-x).exp * (s * x ^ (s - 1)) : ℝ → ℂ) volume 0 Y := +begin + have : (λ x, (-x).exp * (s * x ^ (s - 1)) : ℝ → ℂ) = + (λ x, s * ((-x).exp * x ^ (s - 1)) : ℝ → ℂ), + { ext1, ring, }, + rw [this, interval_integrable_iff_integrable_Ioc_of_le hY], + split, + { refine (continuous_on_const.mul _).ae_strongly_measurable measurable_set_Ioc, + apply (continuous_of_real.comp continuous_neg.exp).continuous_on.mul, + apply continuous_at.continuous_on, + intros x hx, + refine (_ : continuous_at (λ x:ℂ, x ^ (s - 1)) _).comp continuous_of_real.continuous_at, + apply continuous_at_cpow_const, rw of_real_re, exact or.inl hx.1, }, + rw ←has_finite_integral_norm_iff, + simp_rw [norm_eq_abs, map_mul], + refine (((real.Gamma_integral_convergent hs).mono_set + Ioc_subset_Ioi_self).has_finite_integral.congr _).const_mul _, + rw [eventually_eq, ae_restrict_iff'], + { apply ae_of_all, intros x hx, + rw [abs_of_nonneg (exp_pos _).le,abs_cpow_eq_rpow_re_of_pos hx.1], + simp }, + { exact measurable_set_Ioc}, +end + +/-- The recurrence relation for the indefinite version of the `Γ` function. -/ +lemma partial_Gamma_add_one {s : ℂ} (hs: 0 < s.re) {X : ℝ} (hX : 0 ≤ X) : + partial_Gamma (s + 1) X = s * partial_Gamma s X - (-X).exp * X ^ s := +begin + rw [partial_Gamma, partial_Gamma, add_sub_cancel], + have F_der_I: (∀ (x:ℝ), (x ∈ Ioo 0 X) → has_deriv_at (λ x, (-x).exp * x ^ s : ℝ → ℂ) + ( -((-x).exp * x ^ s) + (-x).exp * (s * x ^ (s - 1))) x), + { intros x hx, + have d1 : has_deriv_at (λ (y: ℝ), (-y).exp) (-(-x).exp) x, + { simpa using (has_deriv_at_neg x).exp }, + have d2 : has_deriv_at (λ (y : ℝ), ↑y ^ s) (s * x ^ (s - 1)) x, + { have t := @has_deriv_at.cpow_const _ _ _ s (has_deriv_at_id ↑x) _, + simpa only [mul_one] using t.comp_of_real, + simpa only [id.def, of_real_re, of_real_im, + ne.def, eq_self_iff_true, not_true, or_false, mul_one] using hx.1, }, + simpa only [of_real_neg, neg_mul] using d1.of_real_comp.mul d2 }, + have cont := (continuous_of_real.comp continuous_neg.exp).mul + (continuous_of_real_cpow_const hs), + have der_ible := (Gamma_integrand_deriv_integrable_A hs hX).add + (Gamma_integrand_deriv_integrable_B hs hX), + have int_eval := integral_eq_sub_of_has_deriv_at_of_le hX cont.continuous_on F_der_I der_ible, + -- We are basically done here but manipulating the output into the right form is fiddly. + apply_fun (λ x:ℂ, -x) at int_eval, + rw [interval_integral.integral_add (Gamma_integrand_deriv_integrable_A hs hX) + (Gamma_integrand_deriv_integrable_B hs hX), interval_integral.integral_neg, neg_add, neg_neg] + at int_eval, + rw [eq_sub_of_add_eq int_eval, sub_neg_eq_add, neg_sub, add_comm, add_sub], + simp only [sub_left_inj, add_left_inj], + have : (λ x, (-x).exp * (s * x ^ (s - 1)) : ℝ → ℂ) = (λ x, s * (-x).exp * x ^ (s - 1) : ℝ → ℂ), + { ext1, ring,}, + rw this, + have t := @integral_const_mul 0 X volume _ _ s (λ x:ℝ, (-x).exp * x ^ (s - 1)), + dsimp at t, rw [←t, of_real_zero, zero_cpow], + { rw [mul_zero, add_zero], congr', ext1, ring }, + { contrapose! hs, rw [hs, zero_re] } +end + +/-- The recurrence relation for the `Γ` integral. -/ +theorem Gamma_integral_add_one {s : ℂ} (hs: 0 < s.re) : + Gamma_integral (s + 1) = s * Gamma_integral s := +begin + suffices : tendsto (s+1).partial_Gamma at_top (𝓝 $ s * Gamma_integral s), + { refine tendsto_nhds_unique _ this, + apply tendsto_partial_Gamma, rw [add_re, one_re], linarith, }, + have : (λ X:ℝ, s * partial_Gamma s X - X ^ s * (-X).exp) =ᶠ[at_top] (s+1).partial_Gamma, + { apply eventually_eq_of_mem (Ici_mem_at_top (0:ℝ)), + intros X hX, + rw partial_Gamma_add_one hs (mem_Ici.mp hX), + ring_nf, }, + refine tendsto.congr' this _, + suffices : tendsto (λ X, -X ^ s * (-X).exp : ℝ → ℂ) at_top (𝓝 0), + { simpa using tendsto.add (tendsto.const_mul s (tendsto_partial_Gamma hs)) this }, + rw tendsto_zero_iff_norm_tendsto_zero, + have : (λ (e : ℝ), ‖-(e:ℂ) ^ s * (-e).exp‖ ) =ᶠ[at_top] (λ (e : ℝ), e ^ s.re * (-1 * e).exp ), + { refine eventually_eq_of_mem (Ioi_mem_at_top 0) _, + intros x hx, dsimp only, + rw [norm_eq_abs, map_mul, abs.map_neg, abs_cpow_eq_rpow_re_of_pos hx, + abs_of_nonneg (exp_pos(-x)).le, neg_mul, one_mul],}, + exact (tendsto_congr' this).mpr (tendsto_rpow_mul_exp_neg_mul_at_top_nhds_0 _ _ zero_lt_one), +end + +end Gamma_recurrence + +/-! Now we define `Γ(s)` on the whole complex plane, by recursion. -/ + +section Gamma_def + +/-- The `n`th function in this family is `Γ(s)` if `-n < s.re`, and junk otherwise. -/ +noncomputable def Gamma_aux : ℕ → (ℂ → ℂ) +| 0 := Gamma_integral +| (n+1) := λ s:ℂ, (Gamma_aux n (s+1)) / s + +lemma Gamma_aux_recurrence1 (s : ℂ) (n : ℕ) (h1 : -s.re < ↑n) : + Gamma_aux n s = Gamma_aux n (s+1) / s := +begin + induction n with n hn generalizing s, + { simp only [nat.cast_zero, neg_lt_zero] at h1, + dsimp only [Gamma_aux], rw Gamma_integral_add_one h1, + rw [mul_comm, mul_div_cancel], contrapose! h1, rw h1, + simp }, + { dsimp only [Gamma_aux], + have hh1 : -(s+1).re < n, + { rw [nat.succ_eq_add_one, nat.cast_add, nat.cast_one] at h1, + rw [add_re, one_re], linarith, }, + rw ←(hn (s+1) hh1) } +end + +lemma Gamma_aux_recurrence2 (s : ℂ) (n : ℕ) (h1 : -s.re < ↑n) : + Gamma_aux n s = Gamma_aux (n+1) s := +begin + cases n, + { simp only [nat.cast_zero, neg_lt_zero] at h1, + dsimp only [Gamma_aux], + rw [Gamma_integral_add_one h1, mul_div_cancel_left], + rintro rfl, + rw [zero_re] at h1, + exact h1.false }, + { dsimp only [Gamma_aux], + have : (Gamma_aux n (s + 1 + 1)) / (s+1) = Gamma_aux n (s + 1), + { have hh1 : -(s+1).re < n, + { rw [nat.succ_eq_add_one, nat.cast_add, nat.cast_one] at h1, + rw [add_re, one_re], linarith, }, + rw Gamma_aux_recurrence1 (s+1) n hh1, }, + rw this }, +end + +/-- The `Γ` function (of a complex variable `s`). -/ +@[pp_nodot] def Gamma (s : ℂ) : ℂ := Gamma_aux ⌊1 - s.re⌋₊ s + +lemma Gamma_eq_Gamma_aux (s : ℂ) (n : ℕ) (h1 : -s.re < ↑n) : Gamma s = Gamma_aux n s := +begin + have u : ∀ (k : ℕ), Gamma_aux (⌊1 - s.re⌋₊ + k) s = Gamma s, + { intro k, induction k with k hk, + { simp [Gamma],}, + { rw [←hk, nat.succ_eq_add_one, ←add_assoc], + refine (Gamma_aux_recurrence2 s (⌊1 - s.re⌋₊ + k) _).symm, + rw nat.cast_add, + have i0 := nat.sub_one_lt_floor (1 - s.re), + simp only [sub_sub_cancel_left] at i0, + refine lt_add_of_lt_of_nonneg i0 _, + rw [←nat.cast_zero, nat.cast_le], exact nat.zero_le k, } }, + convert (u $ n - ⌊1 - s.re⌋₊).symm, rw nat.add_sub_of_le, + by_cases (0 ≤ 1 - s.re), + { apply nat.le_of_lt_succ, + exact_mod_cast lt_of_le_of_lt (nat.floor_le h) (by linarith : 1 - s.re < n + 1) }, + { rw nat.floor_of_nonpos, linarith, linarith }, +end + +/-- The recurrence relation for the `Γ` function. -/ +theorem Gamma_add_one (s : ℂ) (h2 : s ≠ 0) : Gamma (s+1) = s * Gamma s := +begin + let n := ⌊1 - s.re⌋₊, + have t1 : -s.re < n, + { simpa only [sub_sub_cancel_left] using nat.sub_one_lt_floor (1 - s.re) }, + have t2 : -(s+1).re < n, + { rw [add_re, one_re], linarith, }, + rw [Gamma_eq_Gamma_aux s n t1, Gamma_eq_Gamma_aux (s+1) n t2, Gamma_aux_recurrence1 s n t1], + field_simp, ring, +end + +theorem Gamma_eq_integral {s : ℂ} (hs : 0 < s.re) : Gamma s = Gamma_integral s := +Gamma_eq_Gamma_aux s 0 (by { norm_cast, linarith }) + +lemma Gamma_one : Gamma 1 = 1 := +by { rw Gamma_eq_integral, simpa using Gamma_integral_one, simp } + +theorem Gamma_nat_eq_factorial (n : ℕ) : Gamma (n+1) = n! := +begin + induction n with n hn, + { simpa using Gamma_one }, + { rw (Gamma_add_one n.succ $ nat.cast_ne_zero.mpr $ nat.succ_ne_zero n), + simp only [nat.cast_succ, nat.factorial_succ, nat.cast_mul], congr, exact hn }, +end + +/-- At `0` the Gamma function is undefined; by convention we assign it the value `0`. -/ +lemma Gamma_zero : Gamma 0 = 0 := +by simp_rw [Gamma, zero_re, sub_zero, nat.floor_one, Gamma_aux, div_zero] + +/-- At `-n` for `n ∈ ℕ`, the Gamma function is undefined; by convention we assign it the value 0. -/ +lemma Gamma_neg_nat_eq_zero (n : ℕ) : Gamma (-n) = 0 := +begin + induction n with n IH, + { rw [nat.cast_zero, neg_zero, Gamma_zero] }, + { have A : -(n.succ : ℂ) ≠ 0, + { rw [neg_ne_zero, nat.cast_ne_zero], + apply nat.succ_ne_zero }, + have : -(n:ℂ) = -↑n.succ + 1, by simp, + rw [this, Gamma_add_one _ A] at IH, + contrapose! IH, + exact mul_ne_zero A IH } +end + +lemma Gamma_conj (s : ℂ) : Gamma (conj s) = conj (Gamma s) := +begin + suffices : ∀ (n:ℕ) (s:ℂ) , Gamma_aux n (conj s) = conj (Gamma_aux n s), from this _ _, + intro n, + induction n with n IH, + { rw Gamma_aux, exact Gamma_integral_conj, }, + { intro s, + rw Gamma_aux, + dsimp only, + rw [div_eq_mul_inv _ s, ring_hom.map_mul, conj_inv, ←div_eq_mul_inv], + suffices : conj s + 1 = conj (s + 1), by rw [this, IH], + rw [ring_hom.map_add, ring_hom.map_one] } +end + +end Gamma_def + +end complex + +/-! Now check that the `Γ` function is differentiable, wherever this makes sense. -/ + +section Gamma_has_deriv + +/-- Integrand for the derivative of the `Γ` function -/ +def dGamma_integrand (s : ℂ) (x : ℝ) : ℂ := exp (-x) * log x * x ^ (s - 1) + +/-- Integrand for the absolute value of the derivative of the `Γ` function -/ +def dGamma_integrand_real (s x : ℝ) : ℝ := |exp (-x) * log x * x ^ (s - 1)| + +lemma dGamma_integrand_is_o_at_top (s : ℝ) : + (λ x : ℝ, exp (-x) * log x * x ^ (s - 1)) =o[at_top] (λ x, exp (-(1/2) * x)) := +begin + refine is_o_of_tendsto (λ x hx, _) _, + { exfalso, exact (-(1/2) * x).exp_pos.ne' hx, }, + have : eventually_eq at_top (λ (x : ℝ), exp (-x) * log x * x ^ (s - 1) / exp (-(1 / 2) * x)) + (λ (x : ℝ), (λ z:ℝ, exp (1 / 2 * z) / z ^ s) x * (λ z:ℝ, z / log z) x)⁻¹, + { refine eventually_of_mem (Ioi_mem_at_top 1) _, + intros x hx, dsimp, + replace hx := lt_trans zero_lt_one (mem_Ioi.mp hx), + rw [real.exp_neg, neg_mul, real.exp_neg, rpow_sub hx], + have : exp x = exp(x/2) * exp(x/2), + { rw [←real.exp_add, add_halves], }, + rw this, field_simp [hx.ne', exp_ne_zero (x/2)], ring, }, + refine tendsto.congr' this.symm (tendsto.inv_tendsto_at_top _), + apply tendsto.at_top_mul_at_top (tendsto_exp_mul_div_rpow_at_top s (1/2) one_half_pos), + refine tendsto.congr' _ ((tendsto_exp_div_pow_at_top 1).comp tendsto_log_at_top), + apply eventually_eq_of_mem (Ioi_mem_at_top (0:ℝ)), + intros x hx, simp [exp_log hx], +end + +/-- Absolute convergence of the integral which will give the derivative of the `Γ` function on +`1 < re s`. -/ +lemma dGamma_integral_abs_convergent (s : ℝ) (hs : 1 < s) : + integrable_on (λ x:ℝ, ‖exp (-x) * log x * x ^ (s-1)‖) (Ioi 0) := +begin + rw [←Ioc_union_Ioi_eq_Ioi (@zero_le_one ℝ _ _ _ _), integrable_on_union], + refine ⟨⟨_, _⟩, _⟩, + { refine continuous_on.ae_strongly_measurable (continuous_on.mul _ _).norm measurable_set_Ioc, + { refine (continuous_exp.comp continuous_neg).continuous_on.mul (continuous_on_log.mono _), + simp, }, + { apply continuous_on_id.rpow_const, intros x hx, right, linarith }, }, + { apply has_finite_integral_of_bounded, + swap, { exact 1 / (s - 1), }, + refine (ae_restrict_iff' measurable_set_Ioc).mpr (ae_of_all _ (λ x hx, _)), + rw [norm_norm, norm_eq_abs, mul_assoc, abs_mul, ←one_mul (1 / (s - 1))], + refine mul_le_mul _ _ (abs_nonneg _) zero_le_one, + { rw [abs_of_pos (exp_pos(-x)), exp_le_one_iff, neg_le, neg_zero], exact hx.1.le }, + { exact (abs_log_mul_self_rpow_lt x (s-1) hx.1 hx.2 (sub_pos.mpr hs)).le }, }, + { have := (dGamma_integrand_is_o_at_top s).is_O.norm_left, + refine integrable_of_is_O_exp_neg one_half_pos (continuous_on.mul _ _).norm this, + { refine (continuous_exp.comp continuous_neg).continuous_on.mul (continuous_on_log.mono _), + simp, }, + { apply continuous_at.continuous_on (λ x hx, _), + apply continuous_at_id.rpow continuous_at_const, + dsimp, right, linarith, }, } +end + +/-- A uniform bound for the `s`-derivative of the `Γ` integrand for `s` in vertical strips. -/ +lemma loc_unif_bound_dGamma_integrand {t : ℂ} {s1 s2 x : ℝ} (ht1 : s1 ≤ t.re) + (ht2: t.re ≤ s2) (hx : 0 < x) : + ‖dGamma_integrand t x‖ ≤ dGamma_integrand_real s1 x + dGamma_integrand_real s2 x := +begin + rcases le_or_lt 1 x with h|h, + { -- case 1 ≤ x + refine le_add_of_nonneg_of_le (abs_nonneg _) _, + rw [dGamma_integrand, dGamma_integrand_real, complex.norm_eq_abs, map_mul, abs_mul, + ←complex.of_real_mul, complex.abs_of_real], + refine mul_le_mul_of_nonneg_left _ (abs_nonneg _), + rw complex.abs_cpow_eq_rpow_re_of_pos hx, + refine le_trans _ (le_abs_self _), + apply rpow_le_rpow_of_exponent_le h, + rw [complex.sub_re, complex.one_re], linarith, }, + { refine le_add_of_le_of_nonneg _ (abs_nonneg _), + rw [dGamma_integrand, dGamma_integrand_real, complex.norm_eq_abs, map_mul, abs_mul, + ←complex.of_real_mul, complex.abs_of_real], + refine mul_le_mul_of_nonneg_left _ (abs_nonneg _), + rw complex.abs_cpow_eq_rpow_re_of_pos hx, + refine le_trans _ (le_abs_self _), + apply rpow_le_rpow_of_exponent_ge hx h.le, + rw [complex.sub_re, complex.one_re], linarith, }, +end + +namespace complex + +/-- The derivative of the `Γ` integral, at any `s ∈ ℂ` with `1 < re s`, is given by the integral +of `exp (-x) * log x * x ^ (s - 1)` over `[0, ∞)`. -/ +theorem has_deriv_at_Gamma_integral {s : ℂ} (hs : 1 < s.re) : + (integrable_on (λ x, real.exp (-x) * real.log x * x ^ (s - 1) : ℝ → ℂ) (Ioi 0) volume) ∧ + (has_deriv_at Gamma_integral (∫ x:ℝ in Ioi 0, real.exp (-x) * real.log x * x ^ (s - 1)) s) := +begin + let ε := (s.re - 1) / 2, + let μ := volume.restrict (Ioi (0:ℝ)), + let bound := (λ x:ℝ, dGamma_integrand_real (s.re - ε) x + dGamma_integrand_real (s.re + ε) x), + have cont : ∀ (t : ℂ), continuous_on (λ x, real.exp (-x) * x ^ (t - 1) : ℝ → ℂ) (Ioi 0), + { intro t, apply (continuous_of_real.comp continuous_neg.exp).continuous_on.mul, + apply continuous_at.continuous_on, intros x hx, + refine (continuous_at_cpow_const _).comp continuous_of_real.continuous_at, + exact or.inl hx, }, + have eps_pos: 0 < ε := div_pos (sub_pos.mpr hs) zero_lt_two, + have hF_meas : ∀ᶠ (t : ℂ) in 𝓝 s, + ae_strongly_measurable (λ x, real.exp(-x) * x ^ (t - 1) : ℝ → ℂ) μ, + { apply eventually_of_forall, intro t, + exact (cont t).ae_strongly_measurable measurable_set_Ioi, }, + have hF'_meas : ae_strongly_measurable (dGamma_integrand s) μ, + { refine continuous_on.ae_strongly_measurable _ measurable_set_Ioi, + have : dGamma_integrand s = (λ x, real.exp (-x) * x ^ (s - 1) * real.log x : ℝ → ℂ), + { ext1, simp only [dGamma_integrand], ring }, + rw this, + refine continuous_on.mul (cont s) (continuous_at.continuous_on _), + exact λ x hx, continuous_of_real.continuous_at.comp (continuous_at_log (mem_Ioi.mp hx).ne'), }, + have h_bound : ∀ᵐ (x : ℝ) ∂μ, ∀ (t : ℂ), t ∈ metric.ball s ε → ‖ dGamma_integrand t x ‖ ≤ bound x, + { refine (ae_restrict_iff' measurable_set_Ioi).mpr (ae_of_all _ (λ x hx, _)), + intros t ht, + rw [metric.mem_ball, complex.dist_eq] at ht, + replace ht := lt_of_le_of_lt (complex.abs_re_le_abs $ t - s ) ht, + rw [complex.sub_re, @abs_sub_lt_iff ℝ _ t.re s.re ((s.re - 1) / 2) ] at ht, + refine loc_unif_bound_dGamma_integrand _ _ hx, + all_goals { simp only [ε], linarith } }, + have bound_integrable : integrable bound μ, + { apply integrable.add, + { refine dGamma_integral_abs_convergent (s.re - ε) _, + field_simp, rw one_lt_div, + { linarith }, { exact zero_lt_two }, }, + { refine dGamma_integral_abs_convergent (s.re + ε) _, linarith, }, }, + have h_diff : ∀ᵐ (x : ℝ) ∂μ, ∀ (t : ℂ), t ∈ metric.ball s ε + → has_deriv_at (λ u, real.exp (-x) * x ^ (u - 1) : ℂ → ℂ) (dGamma_integrand t x) t, + { refine (ae_restrict_iff' measurable_set_Ioi).mpr (ae_of_all _ (λ x hx, _)), + intros t ht, rw mem_Ioi at hx, + simp only [dGamma_integrand], + rw mul_assoc, + apply has_deriv_at.const_mul, + rw [of_real_log hx.le, mul_comm], + have := ((has_deriv_at_id t).sub_const 1).const_cpow (or.inl (of_real_ne_zero.mpr hx.ne')), + rwa mul_one at this }, + exact (has_deriv_at_integral_of_dominated_loc_of_deriv_le eps_pos hF_meas + (Gamma_integral_convergent (zero_lt_one.trans hs)) hF'_meas h_bound bound_integrable h_diff), +end + +lemma differentiable_at_Gamma_aux (s : ℂ) (n : ℕ) (h1 : (1 - s.re) < n ) (h2 : ∀ m : ℕ, s ≠ -m) : + differentiable_at ℂ (Gamma_aux n) s := +begin + induction n with n hn generalizing s, + { refine (has_deriv_at_Gamma_integral _).2.differentiable_at, + rw nat.cast_zero at h1, linarith }, + { dsimp only [Gamma_aux], + specialize hn (s + 1), + have a : 1 - (s + 1).re < ↑n, + { rw nat.cast_succ at h1, rw [complex.add_re, complex.one_re], linarith }, + have b : ∀ m : ℕ, s + 1 ≠ -m, + { intro m, have := h2 (1 + m), + contrapose! this, + rw ←eq_sub_iff_add_eq at this, + simpa using this }, + refine differentiable_at.div (differentiable_at.comp _ (hn a b) _) _ _, + simp, simp, simpa using h2 0 } +end + +theorem differentiable_at_Gamma (s : ℂ) (hs : ∀ m : ℕ, s ≠ -m) : differentiable_at ℂ Gamma s := +begin + let n := ⌊1 - s.re⌋₊ + 1, + have hn : 1 - s.re < n := by exact_mod_cast nat.lt_floor_add_one (1 - s.re), + apply (differentiable_at_Gamma_aux s n hn hs).congr_of_eventually_eq, + let S := { t : ℂ | 1 - t.re < n }, + have : S ∈ 𝓝 s, + { rw mem_nhds_iff, use S, + refine ⟨subset.rfl, _, hn⟩, + have : S = re⁻¹' Ioi (1 - n : ℝ), + { ext, rw [preimage,Ioi, mem_set_of_eq, mem_set_of_eq, mem_set_of_eq], exact sub_lt_comm }, + rw this, + refine continuous.is_open_preimage continuous_re _ is_open_Ioi, }, + apply eventually_eq_of_mem this, + intros t ht, rw mem_set_of_eq at ht, + apply Gamma_eq_Gamma_aux, linarith, +end + +end complex + +end Gamma_has_deriv + +namespace real + +/-- The `Γ` function (of a real variable `s`). -/ +@[pp_nodot] def Gamma (s : ℝ) : ℝ := (complex.Gamma s).re + +lemma Gamma_eq_integral {s : ℝ} (hs : 0 < s) : Gamma s = ∫ x in Ioi 0, exp (-x) * x ^ (s - 1) := +begin + rw [Gamma, complex.Gamma_eq_integral (by rwa complex.of_real_re : 0 < complex.re s)], + dsimp only [complex.Gamma_integral], + simp_rw [←complex.of_real_one, ←complex.of_real_sub], + suffices : ∫ (x : ℝ) in Ioi 0, ↑(exp (-x)) * (x : ℂ) ^ ((s - 1 : ℝ) : ℂ) = + ∫ (x : ℝ) in Ioi 0, ((exp (-x) * x ^ (s - 1) : ℝ) : ℂ), + { rw [this, _root_.integral_of_real, complex.of_real_re], }, + refine set_integral_congr measurable_set_Ioi (λ x hx, _), + push_cast, + rw complex.of_real_cpow (le_of_lt hx), + push_cast, +end + +lemma Gamma_add_one {s : ℝ} (hs : s ≠ 0) : Gamma (s + 1) = s * Gamma s := +begin + simp_rw Gamma, + rw [complex.of_real_add, complex.of_real_one, complex.Gamma_add_one, complex.of_real_mul_re], + rwa complex.of_real_ne_zero, +end + +lemma Gamma_one : Gamma 1 = 1 := +by rw [Gamma, complex.of_real_one, complex.Gamma_one, complex.one_re] + +lemma _root_.complex.Gamma_of_real (s : ℝ) : complex.Gamma (s : ℂ) = Gamma s := +by rw [Gamma, eq_comm, ←complex.conj_eq_iff_re, ←complex.Gamma_conj, complex.conj_of_real] + +theorem Gamma_nat_eq_factorial (n : ℕ) : Gamma (n + 1) = n! := +by rw [Gamma, complex.of_real_add, complex.of_real_nat_cast, complex.of_real_one, + complex.Gamma_nat_eq_factorial, ←complex.of_real_nat_cast, complex.of_real_re] + +/-- At `0` the Gamma function is undefined; by convention we assign it the value `0`. -/ +lemma Gamma_zero : Gamma 0 = 0 := +by simpa only [←complex.of_real_zero, complex.Gamma_of_real, complex.of_real_inj] + using complex.Gamma_zero + +/-- At `-n` for `n ∈ ℕ`, the Gamma function is undefined; by convention we assign it the value `0`. +-/ +lemma Gamma_neg_nat_eq_zero (n : ℕ) : Gamma (-n) = 0 := +begin + simpa only [←complex.of_real_nat_cast, ←complex.of_real_neg, complex.Gamma_of_real, + complex.of_real_eq_zero] using complex.Gamma_neg_nat_eq_zero n, +end + +lemma Gamma_pos_of_pos {s : ℝ} (hs : 0 < s) : 0 < Gamma s := +begin + rw Gamma_eq_integral hs, + have : function.support (λ (x : ℝ), exp (-x) * x ^ (s - 1)) ∩ Ioi 0 = Ioi 0, + { rw inter_eq_right_iff_subset, + intros x hx, + rw function.mem_support, + exact mul_ne_zero (exp_pos _).ne' (rpow_pos_of_pos hx _).ne' }, + rw set_integral_pos_iff_support_of_nonneg_ae, + { rw [this, volume_Ioi, ←ennreal.of_real_zero], + exact ennreal.of_real_lt_top }, + { refine eventually_of_mem (self_mem_ae_restrict measurable_set_Ioi) _, + exact λ x hx, (mul_pos (exp_pos _) (rpow_pos_of_pos hx _)).le }, + { exact Gamma_integral_convergent hs }, +end + +/-- The Gamma function does not vanish on `ℝ` (except at non-positive integers, where the function +is mathematically undefined and we set it to `0` by convention). -/ +lemma Gamma_ne_zero {s : ℝ} (hs : ∀ m : ℕ, s ≠ -m) : Gamma s ≠ 0 := +begin + suffices : ∀ {n : ℕ}, (-(n:ℝ) < s) → Gamma s ≠ 0, + { apply this, + swap, use (⌊-s⌋₊ + 1), + rw [neg_lt, nat.cast_add, nat.cast_one], + exact nat.lt_floor_add_one _ }, + intro n, + induction n generalizing s, + { intro hs, + refine (Gamma_pos_of_pos _).ne', + rwa [nat.cast_zero, neg_zero] at hs }, + { intro hs', + have : Gamma (s + 1) ≠ 0, + { apply n_ih, + { intro m, + specialize hs (1 + m), + contrapose! hs, + rw ←eq_sub_iff_add_eq at hs, + rw hs, + push_cast, + ring }, + { rw [nat.succ_eq_add_one, nat.cast_add, nat.cast_one, neg_add] at hs', + linarith } }, + rw [Gamma_add_one, mul_ne_zero_iff] at this, + { exact this.2 }, + { simpa using hs 0 } }, +end + +lemma Gamma_eq_zero_iff (s : ℝ) : Gamma s = 0 ↔ ∃ m : ℕ, s = -m := +⟨by { contrapose!, exact Gamma_ne_zero }, by { rintro ⟨m, rfl⟩, exact Gamma_neg_nat_eq_zero m }⟩ + +lemma differentiable_at_Gamma {s : ℝ} (hs : ∀ m : ℕ, s ≠ -m) : differentiable_at ℝ Gamma s := +begin + refine ((complex.differentiable_at_Gamma _ _).has_deriv_at).real_of_complex.differentiable_at, + simp_rw [←complex.of_real_nat_cast, ←complex.of_real_neg, ne.def, complex.of_real_inj], + exact hs, +end + +end real diff --git a/src/analysis/special_functions/gamma/beta.lean b/src/analysis/special_functions/gamma/beta.lean new file mode 100644 index 0000000000000..b396e1836867b --- /dev/null +++ b/src/analysis/special_functions/gamma/beta.lean @@ -0,0 +1,509 @@ +/- +Copyright (c) 2023 David Loeffler. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: David Loeffler +-/ +import analysis.convolution +import analysis.special_functions.trigonometric.euler_sine_prod +import analysis.special_functions.gamma.basic + +/-! +# The Beta function, and further properties of the Gamma function + +In this file we define the Beta integral, relate Beta and Gamma functions, and prove some +refined properties of the Gamma function using these relations. + +## Results on the Beta function + +* `complex.beta_integral`: the Beta function `Β(u, v)`, where `u`, `v` are complex with positive + real part. +* `complex.Gamma_mul_Gamma_eq_beta_integral`: the formula + `Gamma u * Gamma v = Gamma (u + v) * beta_integral u v`. + +## Results on the Gamma function + +* `complex.Gamma_ne_zero`: for all `s : ℂ` with `s ∉ {-n : n ∈ ℕ}` we have `Γ s ≠ 0`. +* `complex.Gamma_seq_tendsto_Gamma`: for all `s`, the limit as `n → ∞` of the sequence + `n ↦ n ^ s * n! / (s * (s + 1) * ... * (s + n))` is `Γ(s)`. +* `complex.Gamma_mul_Gamma_one_sub`: Euler's reflection formula + `Gamma s * Gamma (1 - s) = π / sin π s`. +* `real.Gamma_ne_zero`, `real.Gamma_seq_tendsto_Gamma`, + `real.Gamma_mul_Gamma_one_sub`: real versions of the above results. +-/ + +noncomputable theory +open filter interval_integral set real measure_theory +open_locale nat topology big_operators real + +section beta_integral + +/-! ## The Beta function -/ + +namespace complex + +notation `cexp` := complex.exp + +/-- The Beta function `Β (u, v)`, defined as `∫ x:ℝ in 0..1, x ^ (u - 1) * (1 - x) ^ (v - 1)`. -/ +noncomputable def beta_integral (u v : ℂ) : ℂ := +∫ (x:ℝ) in 0..1, x ^ (u - 1) * (1 - x) ^ (v - 1) + +/-- Auxiliary lemma for `beta_integral_convergent`, showing convergence at the left endpoint. -/ +lemma beta_integral_convergent_left {u : ℂ} (hu : 0 < re u) (v : ℂ) : + interval_integrable (λ x, x ^ (u - 1) * (1 - x) ^ (v - 1) : ℝ → ℂ) volume 0 (1 / 2) := +begin + apply interval_integrable.mul_continuous_on, + { refine interval_integral.interval_integrable_cpow' _, + rwa [sub_re, one_re, ←zero_sub, sub_lt_sub_iff_right] }, + { apply continuous_at.continuous_on, + intros x hx, + rw uIcc_of_le (by positivity: (0:ℝ) ≤ 1/2) at hx, + apply continuous_at.cpow, + { exact (continuous_const.sub continuous_of_real).continuous_at }, + { exact continuous_at_const }, + { rw [sub_re, one_re, of_real_re, sub_pos], + exact or.inl (hx.2.trans_lt (by norm_num : (1/2:ℝ) < 1)) } } +end + +/-- The Beta integral is convergent for all `u, v` of positive real part. -/ +lemma beta_integral_convergent {u v : ℂ} (hu : 0 < re u) (hv : 0 < re v) : + interval_integrable (λ x, x ^ (u - 1) * (1 - x) ^ (v - 1) : ℝ → ℂ) volume 0 1 := +begin + refine (beta_integral_convergent_left hu v).trans _, + rw interval_integrable.iff_comp_neg, + convert ((beta_integral_convergent_left hv u).comp_add_right 1).symm, + { ext1 x, + conv_lhs { rw mul_comm }, + congr' 2; + { push_cast, ring } }, + { norm_num }, + { norm_num } +end + +lemma beta_integral_symm (u v : ℂ) : + beta_integral v u = beta_integral u v := +begin + rw [beta_integral, beta_integral], + have := interval_integral.integral_comp_mul_add + (λ x:ℝ, (x:ℂ) ^ (u - 1) * (1 - ↑x) ^ (v - 1)) (neg_one_lt_zero.ne) 1, + rw [inv_neg, inv_one, neg_one_smul, ←interval_integral.integral_symm] at this, + convert this, + { ext1 x, rw mul_comm, congr; + { push_cast, ring } }, + { ring }, { ring } +end + +lemma beta_integral_eval_one_right {u : ℂ} (hu : 0 < re u) : + beta_integral u 1 = 1 / u := +begin + simp_rw [beta_integral, sub_self, cpow_zero, mul_one], + rw integral_cpow (or.inl _), + { rw [of_real_zero, of_real_one, one_cpow, zero_cpow, + sub_zero, sub_add_cancel], + rw sub_add_cancel, + contrapose! hu, rw [hu, zero_re] }, + { rwa [sub_re, one_re, ←sub_pos, sub_neg_eq_add, sub_add_cancel] }, +end + +lemma beta_integral_scaled (s t : ℂ) {a : ℝ} (ha : 0 < a) : + ∫ x in 0..a, (x:ℂ) ^ (s - 1) * (a - x) ^ (t - 1) = a ^ (s + t - 1) * beta_integral s t := +begin + have ha' : (a:ℂ) ≠ 0, from of_real_ne_zero.mpr ha.ne', + rw beta_integral, + have A : (a:ℂ) ^ (s + t - 1) = a * (a ^ (s - 1) * a ^ (t - 1)), + { rw [(by abel : s + t - 1 = 1 + (s - 1) + (t - 1)), + cpow_add _ _ ha', cpow_add 1 _ ha', cpow_one, mul_assoc] }, + rw [A, mul_assoc, ←interval_integral.integral_const_mul ((↑a) ^ _ * _), + ←real_smul, ←(zero_div a), ←div_self ha.ne', + ←interval_integral.integral_comp_div _ ha.ne', zero_div], + simp_rw interval_integral.integral_of_le ha.le, + refine set_integral_congr measurable_set_Ioc (λ x hx, _), + dsimp only, + rw mul_mul_mul_comm, + congr' 1, + { rw [←mul_cpow_of_real_nonneg ha.le (div_pos hx.1 ha).le, of_real_div, mul_div_cancel' _ ha'] }, + { rw [(by push_cast : (1:ℂ) - ↑(x / a) = ↑(1 - x / a)), + ←mul_cpow_of_real_nonneg ha.le (sub_nonneg.mpr $ (div_le_one ha).mpr hx.2)], + push_cast, + rw [mul_sub, mul_one, mul_div_cancel' _ ha'] } +end + +/-- Relation between Beta integral and Gamma function. -/ +lemma Gamma_mul_Gamma_eq_beta_integral {s t : ℂ} (hs : 0 < re s) (ht : 0 < re t) : + Gamma s * Gamma t = Gamma (s + t) * beta_integral s t := +begin + -- Note that we haven't proved (yet) that the Gamma function has no zeroes, so we can't formulate + -- this as a formula for the Beta function. + have conv_int := integral_pos_convolution (Gamma_integral_convergent hs) + (Gamma_integral_convergent ht) (continuous_linear_map.mul ℝ ℂ), + simp_rw continuous_linear_map.mul_apply' at conv_int, + have hst : 0 < re (s + t), + { rw add_re, exact add_pos hs ht }, + rw [Gamma_eq_integral hs, Gamma_eq_integral ht, Gamma_eq_integral hst, Gamma_integral, + Gamma_integral, Gamma_integral, ←conv_int, ←integral_mul_right (beta_integral _ _)], + refine set_integral_congr measurable_set_Ioi (λ x hx, _), + dsimp only, + rw [mul_assoc, ←beta_integral_scaled s t hx, ←interval_integral.integral_const_mul], + congr' 1 with y:1, + push_cast, + suffices : cexp (-x) = cexp (-y) * cexp (-(x - y)), + { rw this, ring }, + { rw ←complex.exp_add, congr' 1, abel }, +end + +/-- Recurrence formula for the Beta function. -/ +lemma beta_integral_recurrence {u v : ℂ} (hu : 0 < re u) (hv : 0 < re v) : + u * beta_integral u (v + 1) = v * beta_integral (u + 1) v := +begin + -- NB: If we knew `Gamma (u + v + 1) ≠ 0` this would be an easy consequence of + -- `Gamma_mul_Gamma_eq_beta_integral`; but we don't know that yet. We will prove it later, but + -- this lemma is needed in the proof. So we give a (somewhat laborious) direct argument. + let F : ℝ → ℂ := λ x, x ^ u * (1 - x) ^ v, + have hu' : 0 < re (u + 1), by { rw [add_re, one_re], positivity }, + have hv' : 0 < re (v + 1), by { rw [add_re, one_re], positivity }, + have hc : continuous_on F (Icc 0 1), + { refine (continuous_at.continuous_on (λ x hx, _)).mul (continuous_at.continuous_on (λ x hx, _)), + { refine (continuous_at_cpow_const_of_re_pos (or.inl _) hu).comp + continuous_of_real.continuous_at, + rw of_real_re, exact hx.1 }, + { refine (continuous_at_cpow_const_of_re_pos (or.inl _) hv).comp + (continuous_const.sub continuous_of_real).continuous_at, + rw [sub_re, one_re, of_real_re, sub_nonneg], + exact hx.2 } }, + have hder : ∀ (x : ℝ), x ∈ Ioo (0:ℝ) 1 → has_deriv_at F + (u * (↑x ^ (u - 1) * (1 - ↑x) ^ v) - v * (↑x ^ u * (1 - ↑x) ^ (v - 1))) x, + { intros x hx, + have U : has_deriv_at (λ y:ℂ, y ^ u) (u * ↑x ^ (u - 1)) ↑x, + { have := has_deriv_at.cpow_const (has_deriv_at_id ↑x) (or.inl _), + { rw mul_one at this, exact this }, + { rw [id.def, of_real_re], exact hx.1 } }, + have V : has_deriv_at (λ y:ℂ, (1 - y) ^ v) (-v * (1 - ↑x) ^ (v - 1)) ↑x, + { have A := has_deriv_at.cpow_const (has_deriv_at_id (1 - ↑x)) (or.inl _), + rotate, { exact v }, + { rw [id.def, sub_re, one_re, of_real_re, sub_pos], exact hx.2 }, + simp_rw [id.def] at A, + have B : has_deriv_at (λ y:ℂ, 1 - y) (-1) ↑x, + { apply has_deriv_at.const_sub, apply has_deriv_at_id }, + convert has_deriv_at.comp ↑x A B using 1, + ring }, + convert (U.mul V).comp_of_real, + ring }, + have h_int := ((beta_integral_convergent hu hv').const_mul u).sub + ((beta_integral_convergent hu' hv).const_mul v), + dsimp only at h_int, + rw [add_sub_cancel, add_sub_cancel] at h_int, + have int_ev := interval_integral.integral_eq_sub_of_has_deriv_at_of_le zero_le_one hc hder h_int, + have hF0 : F 0 = 0, + { simp only [mul_eq_zero, of_real_zero, cpow_eq_zero_iff, eq_self_iff_true, + ne.def, true_and, sub_zero, one_cpow, one_ne_zero, or_false], + contrapose! hu, rw [hu, zero_re] }, + have hF1 : F 1 = 0, + { simp only [mul_eq_zero, of_real_one, one_cpow, one_ne_zero, sub_self, + cpow_eq_zero_iff, eq_self_iff_true, ne.def, true_and, false_or], + contrapose! hv, rw [hv, zero_re] }, + rw [hF0, hF1, sub_zero, interval_integral.integral_sub, + interval_integral.integral_const_mul, interval_integral.integral_const_mul] at int_ev, + { rw [beta_integral, beta_integral, ←sub_eq_zero], + convert int_ev; + { ext1 x, congr, abel } }, + { apply interval_integrable.const_mul, + convert beta_integral_convergent hu hv', + ext1 x, rw add_sub_cancel }, + { apply interval_integrable.const_mul, + convert beta_integral_convergent hu' hv, + ext1 x, rw add_sub_cancel }, +end + +/-- Explicit formula for the Beta function when second argument is a positive integer. -/ +lemma beta_integral_eval_nat_add_one_right {u : ℂ} (hu : 0 < re u) (n : ℕ) : + beta_integral u (n + 1) = n! / ∏ (j:ℕ) in finset.range (n + 1), (u + j) := +begin + induction n with n IH generalizing u, + { rw [nat.cast_zero, zero_add, beta_integral_eval_one_right hu, + nat.factorial_zero, nat.cast_one, zero_add, finset.prod_range_one, nat.cast_zero, add_zero] }, + { have := beta_integral_recurrence hu (_ : 0 < re n.succ), + swap, { rw [←of_real_nat_cast, of_real_re], positivity }, + rw [mul_comm u _, ←eq_div_iff] at this, + swap, { contrapose! hu, rw [hu, zero_re] }, + rw [this, finset.prod_range_succ', nat.cast_succ, IH], + swap, { rw [add_re, one_re], positivity }, + rw [nat.factorial_succ, nat.cast_mul, nat.cast_add, nat.cast_one, nat.cast_zero, add_zero, + ←mul_div_assoc, ←div_div], + congr' 3 with j:1, + push_cast, abel } +end + +end complex + +end beta_integral + +section limit_formula + +/-! ## The Euler limit formula -/ + +namespace complex + +/-- The sequence with `n`-th term `n ^ s * n! / (s * (s + 1) * ... * (s + n))`, for complex `s`. +We will show that this tends to `Γ(s)` as `n → ∞`. -/ +noncomputable def Gamma_seq (s : ℂ) (n : ℕ) := +(n:ℂ) ^ s * n! / ∏ (j:ℕ) in finset.range (n + 1), (s + j) + +lemma Gamma_seq_eq_beta_integral_of_re_pos {s : ℂ} (hs : 0 < re s) (n : ℕ) : + Gamma_seq s n = n ^ s * beta_integral s (n + 1) := +by rw [Gamma_seq, beta_integral_eval_nat_add_one_right hs n, ←mul_div_assoc] + +lemma Gamma_seq_add_one_left (s : ℂ) {n : ℕ} (hn : n ≠ 0) : + (Gamma_seq (s + 1) n) / s = n / (n + 1 + s) * Gamma_seq s n := +begin + conv_lhs { rw [Gamma_seq, finset.prod_range_succ, div_div] }, + conv_rhs { rw [Gamma_seq, finset.prod_range_succ', nat.cast_zero, add_zero, div_mul_div_comm, + ←mul_assoc, ←mul_assoc, mul_comm _ (finset.prod _ _)] }, + congr' 3, + { rw [cpow_add _ _ (nat.cast_ne_zero.mpr hn), cpow_one, mul_comm] }, + { refine finset.prod_congr (by refl) (λ x hx, _), + push_cast, ring }, + { abel } +end + +lemma Gamma_seq_eq_approx_Gamma_integral {s : ℂ} (hs : 0 < re s) {n : ℕ} (hn : n ≠ 0) : + Gamma_seq s n = ∫ x:ℝ in 0..n, ↑((1 - x / n) ^ n) * (x:ℂ) ^ (s - 1) := +begin + have : ∀ (x : ℝ), x = x / n * n, by { intro x, rw div_mul_cancel, exact nat.cast_ne_zero.mpr hn }, + conv in (↑_ ^ _) { congr, rw this x }, + rw Gamma_seq_eq_beta_integral_of_re_pos hs, + rw [beta_integral, @interval_integral.integral_comp_div _ _ _ _ 0 n _ + (λ x, ↑((1 - x) ^ n) * ↑(x * ↑n) ^ (s - 1) : ℝ → ℂ) (nat.cast_ne_zero.mpr hn), + real_smul, zero_div, div_self, add_sub_cancel, ←interval_integral.integral_const_mul, + ←interval_integral.integral_const_mul], + swap, { exact nat.cast_ne_zero.mpr hn }, + simp_rw interval_integral.integral_of_le zero_le_one, + refine set_integral_congr measurable_set_Ioc (λ x hx, _), + push_cast, + have hn' : (n : ℂ) ≠ 0, from nat.cast_ne_zero.mpr hn, + have A : (n : ℂ) ^ s = (n : ℂ) ^ (s - 1) * n, + { conv_lhs { rw [(by ring : s = (s - 1) + 1), cpow_add _ _ hn'] }, + simp }, + have B : ((x : ℂ) * ↑n) ^ (s - 1) = (x : ℂ) ^ (s - 1) * ↑n ^ (s - 1), + { rw [←of_real_nat_cast, + mul_cpow_of_real_nonneg hx.1.le (nat.cast_pos.mpr (nat.pos_of_ne_zero hn)).le] }, + rw [A, B, cpow_nat_cast], ring, +end + +/-- The main techical lemma for `Gamma_seq_tendsto_Gamma`, expressing the integral defining the +Gamma function for `0 < re s` as the limit of a sequence of integrals over finite intervals. -/ +lemma approx_Gamma_integral_tendsto_Gamma_integral {s : ℂ} (hs : 0 < re s) : + tendsto (λ n:ℕ, ∫ x:ℝ in 0..n, ↑((1 - x / n) ^ n) * (x:ℂ) ^ (s - 1)) at_top (𝓝 $ Gamma s) := +begin + rw [Gamma_eq_integral hs], + -- We apply dominated convergence to the following function, which we will show is uniformly + -- bounded above by the Gamma integrand `exp (-x) * x ^ (re s - 1)`. + let f : ℕ → ℝ → ℂ := λ n, indicator (Ioc 0 (n:ℝ)) + (λ x:ℝ, ↑((1 - x / n) ^ n) * (x:ℂ) ^ (s - 1)), + -- integrability of f + have f_ible : ∀ (n:ℕ), integrable (f n) (volume.restrict (Ioi 0)), + { intro n, + rw [integrable_indicator_iff (measurable_set_Ioc : measurable_set (Ioc (_:ℝ) _)), + integrable_on, measure.restrict_restrict_of_subset Ioc_subset_Ioi_self, ←integrable_on, + ←interval_integrable_iff_integrable_Ioc_of_le (by positivity : (0:ℝ) ≤ n)], + apply interval_integrable.continuous_on_mul, + { refine interval_integral.interval_integrable_cpow' _, + rwa [sub_re, one_re, ←zero_sub, sub_lt_sub_iff_right] }, + { apply continuous.continuous_on, continuity } }, + -- pointwise limit of f + have f_tends : ∀ x:ℝ, x ∈ Ioi (0:ℝ) → + tendsto (λ n:ℕ, f n x) at_top (𝓝 $ ↑(real.exp (-x)) * (x:ℂ) ^ (s - 1)), + { intros x hx, + apply tendsto.congr', + show ∀ᶠ n:ℕ in at_top, ↑((1 - x / n) ^ n) * (x:ℂ) ^ (s - 1) = f n x, + { refine eventually.mp (eventually_ge_at_top ⌈x⌉₊) (eventually_of_forall (λ n hn, _)), + rw nat.ceil_le at hn, + dsimp only [f], + rw indicator_of_mem, + exact ⟨hx, hn⟩ }, + { simp_rw mul_comm _ (↑x ^ _), + refine (tendsto.comp (continuous_of_real.tendsto _) _).const_mul _, + convert tendsto_one_plus_div_pow_exp (-x), + ext1 n, + rw [neg_div, ←sub_eq_add_neg] } }, + -- let `convert` identify the remaining goals + convert tendsto_integral_of_dominated_convergence _ (λ n, (f_ible n).1) + (real.Gamma_integral_convergent hs) _ + ((ae_restrict_iff' measurable_set_Ioi).mpr (ae_of_all _ f_tends)), + -- limit of f is the integrand we want + { ext1 n, + rw [integral_indicator (measurable_set_Ioc : measurable_set (Ioc (_:ℝ) _)), + interval_integral.integral_of_le (by positivity: 0 ≤ (n:ℝ)), + measure.restrict_restrict_of_subset Ioc_subset_Ioi_self] }, + -- f is uniformly bounded by the Gamma integrand + { intro n, + refine (ae_restrict_iff' measurable_set_Ioi).mpr (ae_of_all _ (λ x hx, _)), + dsimp only [f], + rcases lt_or_le (n:ℝ) x with hxn | hxn, + { rw [indicator_of_not_mem (not_mem_Ioc_of_gt hxn), norm_zero, + mul_nonneg_iff_right_nonneg_of_pos (exp_pos _)], + exact rpow_nonneg_of_nonneg (le_of_lt hx) _ }, + { rw [indicator_of_mem (mem_Ioc.mpr ⟨hx, hxn⟩), norm_mul, complex.norm_eq_abs, + complex.abs_of_nonneg + (pow_nonneg (sub_nonneg.mpr $ div_le_one_of_le hxn $ by positivity) _), + complex.norm_eq_abs, abs_cpow_eq_rpow_re_of_pos hx, sub_re, one_re, + mul_le_mul_right (rpow_pos_of_pos hx _ )], + exact one_sub_div_pow_le_exp_neg hxn } } +end + +/-- Euler's limit formula for the complex Gamma function. -/ +lemma Gamma_seq_tendsto_Gamma (s : ℂ) : + tendsto (Gamma_seq s) at_top (𝓝 $ Gamma s) := +begin + suffices : ∀ m : ℕ, (-↑m < re s) → tendsto (Gamma_seq s) at_top (𝓝 $ Gamma_aux m s), + { rw Gamma, + apply this, + rw neg_lt, + rcases lt_or_le 0 (re s) with hs | hs, + { exact (neg_neg_of_pos hs).trans_le (nat.cast_nonneg _), }, + { refine (nat.lt_floor_add_one _).trans_le _, + rw [sub_eq_neg_add, nat.floor_add_one (neg_nonneg.mpr hs), nat.cast_add_one] } }, + intro m, + induction m with m IH generalizing s, + { -- Base case: `0 < re s`, so Gamma is given by the integral formula + intro hs, + rw [nat.cast_zero, neg_zero] at hs, + rw [←Gamma_eq_Gamma_aux], + { refine tendsto.congr' _ (approx_Gamma_integral_tendsto_Gamma_integral hs), + refine (eventually_ne_at_top 0).mp (eventually_of_forall (λ n hn, _)), + exact (Gamma_seq_eq_approx_Gamma_integral hs hn).symm }, + { rwa [nat.cast_zero, neg_lt_zero] } }, + { -- Induction step: use recurrence formulae in `s` for Gamma and Gamma_seq + intro hs, + rw [nat.cast_succ, neg_add, ←sub_eq_add_neg, sub_lt_iff_lt_add, ←one_re, ←add_re] at hs, + rw Gamma_aux, + have := tendsto.congr' ((eventually_ne_at_top 0).mp (eventually_of_forall (λ n hn, _))) + ((IH _ hs).div_const s), + swap 3, { exact Gamma_seq_add_one_left s hn }, -- doesn't work if inlined? + conv at this in (_ / _ * _) { rw mul_comm }, + rwa [←mul_one (Gamma_aux m (s + 1) / s), tendsto_mul_iff_of_ne_zero _ (one_ne_zero' ℂ)] at this, + simp_rw add_assoc, + exact tendsto_coe_nat_div_add_at_top (1 + s) } +end + +end complex + +end limit_formula + +section gamma_reflection +/-! ## The reflection formula -/ + +namespace complex + +lemma Gamma_seq_mul (z : ℂ) {n : ℕ} (hn : n ≠ 0) : + Gamma_seq z n * Gamma_seq (1 - z) n = + n / (n + 1 - z) * (1 / (z * ∏ j in finset.range n, (1 - z ^ 2 / (j + 1) ^ 2))) := +begin + -- also true for n = 0 but we don't need it + have aux : ∀ (a b c d : ℂ), a * b * (c * d) = a * c * (b * d), by { intros, ring }, + rw [Gamma_seq, Gamma_seq, div_mul_div_comm, aux, ←pow_two], + have : (n : ℂ) ^ z * n ^ (1 - z) = n, + { rw [←cpow_add _ _ (nat.cast_ne_zero.mpr hn), add_sub_cancel'_right, cpow_one] }, + rw [this, finset.prod_range_succ', finset.prod_range_succ, aux, ←finset.prod_mul_distrib, + nat.cast_zero, add_zero, add_comm (1 - z) n, ←add_sub_assoc], + have : ∀ (j : ℕ), (z + ↑(j + 1)) * (1 - z + ↑j) = ↑((j + 1) ^ 2) * (1 - z ^ 2 / (↑j + 1) ^ 2), + { intro j, + push_cast, + have : (j:ℂ) + 1 ≠ 0, by { rw [←nat.cast_succ, nat.cast_ne_zero], exact nat.succ_ne_zero j }, + field_simp, ring }, + simp_rw this, + rw [finset.prod_mul_distrib, ←nat.cast_prod, finset.prod_pow, + finset.prod_range_add_one_eq_factorial, nat.cast_pow, + (by {intros, ring} : ∀ (a b c d : ℂ), a * b * (c * d) = a * (d * (b * c))), + ←div_div, mul_div_cancel, ←div_div, mul_comm z _, mul_one_div], + exact pow_ne_zero 2 (nat.cast_ne_zero.mpr $ nat.factorial_ne_zero n), +end + +/-- Euler's reflection formula for the complex Gamma function. -/ +theorem Gamma_mul_Gamma_one_sub (z : ℂ) : Gamma z * Gamma (1 - z) = π / sin (π * z) := +begin + have pi_ne : (π : ℂ) ≠ 0, from complex.of_real_ne_zero.mpr pi_ne_zero, + by_cases hs : sin (↑π * z) = 0, + { -- first deal with silly case z = integer + rw [hs, div_zero], + rw [←neg_eq_zero, ←complex.sin_neg, ←mul_neg, complex.sin_eq_zero_iff, mul_comm] at hs, + obtain ⟨k, hk⟩ := hs, + rw [mul_eq_mul_right_iff, eq_false_intro (of_real_ne_zero.mpr pi_pos.ne'), or_false, + neg_eq_iff_eq_neg] at hk, + rw hk, + cases k, + { rw [int.cast_of_nat, complex.Gamma_neg_nat_eq_zero, zero_mul] }, + { rw [int.cast_neg_succ_of_nat, neg_neg, nat.cast_add, nat.cast_one, add_comm, sub_add_cancel', + complex.Gamma_neg_nat_eq_zero, mul_zero] } }, + refine tendsto_nhds_unique ((Gamma_seq_tendsto_Gamma z).mul (Gamma_seq_tendsto_Gamma $ 1 - z)) _, + have : ↑π / sin (↑π * z) = 1 * (π / sin (π * z)), by rw one_mul, rw this, + refine tendsto.congr' ((eventually_ne_at_top 0).mp + (eventually_of_forall (λ n hn, (Gamma_seq_mul z hn).symm))) (tendsto.mul _ _), + { convert tendsto_coe_nat_div_add_at_top (1 - z), ext1 n, rw add_sub_assoc }, + { have : ↑π / sin (↑π * z) = 1 / (sin (π * z) / π), by field_simp, rw this, + refine tendsto_const_nhds.div _ (div_ne_zero hs pi_ne), + rw [←tendsto_mul_iff_of_ne_zero tendsto_const_nhds pi_ne, div_mul_cancel _ pi_ne], + convert tendsto_euler_sin_prod z, + ext1 n, rw [mul_comm, ←mul_assoc] }, +end + +/-- The Gamma function does not vanish on `ℂ` (except at non-positive integers, where the function +is mathematically undefined and we set it to `0` by convention). -/ +theorem Gamma_ne_zero {s : ℂ} (hs : ∀ m : ℕ, s ≠ -m) : Gamma s ≠ 0 := +begin + by_cases h_im : s.im = 0, + { have : s = ↑s.re, + { conv_lhs { rw ←complex.re_add_im s }, rw [h_im, of_real_zero, zero_mul, add_zero] }, + rw [this, Gamma_of_real, of_real_ne_zero], + refine real.Gamma_ne_zero (λ n, _), + specialize hs n, + contrapose! hs, + rwa [this, ←of_real_nat_cast, ←of_real_neg, of_real_inj] }, + { have : sin (↑π * s) ≠ 0, + { rw complex.sin_ne_zero_iff, + intro k, + apply_fun im, + rw [of_real_mul_im, ←of_real_int_cast, ←of_real_mul, of_real_im], + exact mul_ne_zero real.pi_pos.ne' h_im }, + have A := div_ne_zero (of_real_ne_zero.mpr real.pi_pos.ne') this, + rw [←complex.Gamma_mul_Gamma_one_sub s, mul_ne_zero_iff] at A, + exact A.1 } +end + +lemma Gamma_eq_zero_iff (s : ℂ) : Gamma s = 0 ↔ ∃ m : ℕ, s = -m := +begin + split, + { contrapose!, exact Gamma_ne_zero }, + { rintro ⟨m, rfl⟩, exact Gamma_neg_nat_eq_zero m }, +end + +end complex + +namespace real + +/-- The sequence with `n`-th term `n ^ s * n! / (s * (s + 1) * ... * (s + n))`, for real `s`. We +will show that this tends to `Γ(s)` as `n → ∞`. -/ +noncomputable def Gamma_seq (s : ℝ) (n : ℕ) := +(n : ℝ) ^ s * n! / ∏ (j : ℕ) in finset.range (n + 1), (s + j) + +/-- Euler's limit formula for the real Gamma function. -/ +lemma Gamma_seq_tendsto_Gamma (s : ℝ) : tendsto (Gamma_seq s) at_top (𝓝 $ Gamma s) := +begin + suffices : tendsto (coe ∘ Gamma_seq s : ℕ → ℂ) at_top (𝓝 $ complex.Gamma s), + from (complex.continuous_re.tendsto (complex.Gamma ↑s)).comp this, + convert complex.Gamma_seq_tendsto_Gamma s, + ext1 n, + dsimp only [Gamma_seq, function.comp_app, complex.Gamma_seq], + push_cast, + rw [complex.of_real_cpow n.cast_nonneg, complex.of_real_nat_cast] +end + +/-- Euler's reflection formula for the real Gamma function. -/ +lemma Gamma_mul_Gamma_one_sub (s : ℝ) : Gamma s * Gamma (1 - s) = π / sin (π * s) := +begin + simp_rw [←complex.of_real_inj, complex.of_real_div, complex.of_real_sin, + complex.of_real_mul, ←complex.Gamma_of_real, complex.of_real_sub, complex.of_real_one], + exact complex.Gamma_mul_Gamma_one_sub s +end + +end real + +end gamma_reflection diff --git a/src/analysis/special_functions/gamma/bohr_mollerup.lean b/src/analysis/special_functions/gamma/bohr_mollerup.lean new file mode 100644 index 0000000000000..b6aa7efc4ecbb --- /dev/null +++ b/src/analysis/special_functions/gamma/bohr_mollerup.lean @@ -0,0 +1,380 @@ +/- +Copyright (c) 2023 David Loeffler. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: David Loeffler +-/ +import analysis.special_functions.gamma.basic + +/-! # Convexity properties of the Gamma function + +In this file, we prove that `Gamma` and `log ∘ Gamma` are convex functions on the positive real +line. We then prove the Bohr-Mollerup theorem, which characterises `Gamma` as the *unique* +positive-real-valued, log-convex function on the positive reals satisfying `f (x + 1) = x f x` and +`f 1 = 1`. + +The proof of the Bohr-Mollerup theorem is bound up with the proof of (a weak form of) the Euler +limit formula, `real.bohr_mollerup.tendsto_log_gamma_seq`, stating that for positive +real `x` the sequence `x * log n + log n! - ∑ (m : ℕ) in finset.range (n + 1), log (x + m)` +tends to `log Γ(x)` as `n → ∞`. We prove that any function satisfying the hypotheses of the +Bohr-Mollerup theorem must agree with the limit in the Euler limit formula, so there is at most one +such function; then we show that `Γ` satisfies these conditions. + +Since most of the auxiliary lemmas for the Bohr-Mollerup theorem are of no relevance outside the +context of this proof, we place them in a separate namespace `real.bohr_mollerup` to avoid clutter. +(This includes the logarithmic form of the Euler limit formula, since later we will prove a more +general form of the Euler limit formula valid for any real or complex `x`; see +`real.Gamma_seq_tendsto_Gamma` and `complex.Gamma_seq_tendsto_Gamma` in the file +`analysis.special_functions.gamma.beta`.) +-/ + +noncomputable theory +open filter set measure_theory +open_locale nat ennreal topology big_operators + +namespace real + +section convexity + +/-- Log-convexity of the Gamma function on the positive reals (stated in multiplicative form), +proved using the Hölder inequality applied to Euler's integral. -/ +lemma Gamma_mul_add_mul_le_rpow_Gamma_mul_rpow_Gamma {s t a b : ℝ} + (hs : 0 < s) (ht : 0 < t) (ha : 0 < a) (hb : 0 < b) (hab : a + b = 1) : + Gamma (a * s + b * t) ≤ Gamma s ^ a * Gamma t ^ b := +begin + -- We will apply Hölder's inequality, for the conjugate exponents `p = 1 / a` + -- and `q = 1 / b`, to the functions `f a s` and `f b t`, where `f` is as follows: + let f : ℝ → ℝ → ℝ → ℝ := λ c u x, exp (-c * x) * x ^ (c * (u - 1)), + have e : is_conjugate_exponent (1 / a) (1 / b) := real.is_conjugate_exponent_one_div ha hb hab, + have hab' : b = 1 - a := by linarith, + have hst : 0 < a * s + b * t := add_pos (mul_pos ha hs) (mul_pos hb ht), + -- some properties of f: + have posf : ∀ (c u x : ℝ), x ∈ Ioi (0:ℝ) → 0 ≤ f c u x := + λ c u x hx, mul_nonneg (exp_pos _).le (rpow_pos_of_pos hx _).le, + have posf' : ∀ (c u : ℝ), ∀ᵐ (x : ℝ) ∂volume.restrict (Ioi 0), 0 ≤ f c u x := + λ c u, (ae_restrict_iff' measurable_set_Ioi).mpr (ae_of_all _ (posf c u)), + have fpow : ∀ {c x : ℝ} (hc : 0 < c) (u : ℝ) (hx : 0 < x), + exp (-x) * x ^ (u - 1) = f c u x ^ (1 / c), + { intros c x hc u hx, + dsimp only [f], + rw [mul_rpow (exp_pos _).le ((rpow_nonneg_of_nonneg hx.le) _), ←exp_mul, ←rpow_mul hx.le], + congr' 2; + { field_simp [hc.ne'], ring } }, + -- show `f c u` is in `ℒp` for `p = 1/c`: + have f_mem_Lp : ∀ {c u : ℝ} (hc : 0 < c) (hu : 0 < u), + mem_ℒp (f c u) (ennreal.of_real (1 / c)) (volume.restrict (Ioi 0)), + { intros c u hc hu, + have A : ennreal.of_real (1 / c) ≠ 0, + by rwa [ne.def, ennreal.of_real_eq_zero, not_le, one_div_pos], + have B : ennreal.of_real (1 / c) ≠ ∞, from ennreal.of_real_ne_top, + rw [←mem_ℒp_norm_rpow_iff _ A B, ennreal.to_real_of_real (one_div_nonneg.mpr hc.le), + ennreal.div_self A B, mem_ℒp_one_iff_integrable], + { apply integrable.congr (Gamma_integral_convergent hu), + refine eventually_eq_of_mem (self_mem_ae_restrict measurable_set_Ioi) (λ x hx, _), + dsimp only, + rw fpow hc u hx, + congr' 1, + exact (norm_of_nonneg (posf _ _ x hx)).symm }, + { refine continuous_on.ae_strongly_measurable _ measurable_set_Ioi, + refine (continuous.continuous_on _).mul (continuous_at.continuous_on (λ x hx, _)), + { exact continuous_exp.comp (continuous_const.mul continuous_id'), }, + { exact continuous_at_rpow_const _ _ (or.inl (ne_of_lt hx).symm), } } }, + -- now apply Hölder: + rw [Gamma_eq_integral hs, Gamma_eq_integral ht, Gamma_eq_integral hst], + convert measure_theory.integral_mul_le_Lp_mul_Lq_of_nonneg e (posf' a s) (posf' b t) + (f_mem_Lp ha hs) (f_mem_Lp hb ht) using 1, + { refine set_integral_congr measurable_set_Ioi (λ x hx, _), + dsimp only [f], + have A : exp (-x) = exp (-a * x) * exp (-b * x), + { rw [←exp_add, ←add_mul, ←neg_add, hab, neg_one_mul] }, + have B : x ^ (a * s + b * t - 1) = (x ^ (a * (s - 1))) * (x ^ (b * (t - 1))), + { rw [←rpow_add hx, hab'], congr' 1, ring }, + rw [A, B], + ring }, + { rw [one_div_one_div, one_div_one_div], + congr' 2; + exact set_integral_congr measurable_set_Ioi (λ x hx, fpow (by assumption) _ hx) }, +end + +lemma convex_on_log_Gamma : convex_on ℝ (Ioi 0) (log ∘ Gamma) := +begin + refine convex_on_iff_forall_pos.mpr ⟨convex_Ioi _, λ x hx y hy a b ha hb hab, _⟩, + have : b = 1 - a := by linarith, subst this, + simp_rw [function.comp_app, smul_eq_mul], + rw [←log_rpow (Gamma_pos_of_pos hy), ←log_rpow (Gamma_pos_of_pos hx), + ←log_mul + ((rpow_pos_of_pos (Gamma_pos_of_pos hx) _).ne') (rpow_pos_of_pos (Gamma_pos_of_pos hy) _).ne', + log_le_log + (Gamma_pos_of_pos (add_pos (mul_pos ha hx) (mul_pos hb hy))) + (mul_pos + (rpow_pos_of_pos (Gamma_pos_of_pos hx) _) (rpow_pos_of_pos (Gamma_pos_of_pos hy) _))], + exact Gamma_mul_add_mul_le_rpow_Gamma_mul_rpow_Gamma hx hy ha hb hab, +end + +lemma convex_on_Gamma : convex_on ℝ (Ioi 0) Gamma := +begin + refine ⟨convex_Ioi 0, λ x hx y hy a b ha hb hab, _⟩, + have := convex_on.comp (convex_on_exp.subset (subset_univ _) _) convex_on_log_Gamma + (λ u hu v hv huv, exp_le_exp.mpr huv), + convert this.2 hx hy ha hb hab, + { rw [function.comp_app, exp_log (Gamma_pos_of_pos $ this.1 hx hy ha hb hab)] }, + { rw [function.comp_app, exp_log (Gamma_pos_of_pos hx)] }, + { rw [function.comp_app, exp_log (Gamma_pos_of_pos hy)] }, + { rw convex_iff_is_preconnected, + refine is_preconnected_Ioi.image _ (λ x hx, continuous_at.continuous_within_at _), + refine (differentiable_at_Gamma (λ m, _)).continuous_at.log (Gamma_pos_of_pos hx).ne', + exact (neg_lt_iff_pos_add.mpr (add_pos_of_pos_of_nonneg hx (nat.cast_nonneg m))).ne' } +end + +end convexity + +section bohr_mollerup + +namespace bohr_mollerup + +/-- The function `n ↦ x log n + log n! - (log x + ... + log (x + n))`, which we will show tends to +`log (Gamma x)` as `n → ∞`. -/ +def log_gamma_seq (x : ℝ) (n : ℕ) : ℝ := +x * log n + log n! - ∑ (m : ℕ) in finset.range (n + 1), log (x + m) + +variables {f : ℝ → ℝ} {x : ℝ} {n : ℕ} + +lemma f_nat_eq (hf_feq : ∀ {y:ℝ}, 0 < y → f (y + 1) = f y + log y) (hn : n ≠ 0) : + f n = f 1 + log (n - 1)! := +begin + refine nat.le_induction (by simp) (λ m hm IH, _) n (nat.one_le_iff_ne_zero.2 hn), + have A : 0 < (m : ℝ), from nat.cast_pos.2 hm, + simp only [hf_feq A, nat.cast_add, algebra_map.coe_one, nat.add_succ_sub_one, add_zero], + rw [IH, add_assoc, ← log_mul (nat.cast_ne_zero.mpr (nat.factorial_ne_zero _)) A.ne', + ← nat.cast_mul], + conv_rhs { rw [← nat.succ_pred_eq_of_pos hm, nat.factorial_succ, mul_comm] }, + congr, + exact (nat.succ_pred_eq_of_pos hm).symm +end + +lemma f_add_nat_eq (hf_feq : ∀ {y:ℝ}, 0 < y → f (y + 1) = f y + log y) (hx : 0 < x) (n : ℕ) : + f (x + n) = f x + ∑ (m : ℕ) in finset.range n, log (x + m) := +begin + induction n with n hn, + { simp }, + { have : x + n.succ = (x + n) + 1, + { push_cast, ring }, + rw [this, hf_feq, hn], + rw [finset.range_succ, finset.sum_insert (finset.not_mem_range_self)], + abel, + linarith [(nat.cast_nonneg n : 0 ≤ (n:ℝ))] }, +end + +/-- Linear upper bound for `f (x + n)` on unit interval -/ +lemma f_add_nat_le + (hf_conv : convex_on ℝ (Ioi 0) f) (hf_feq : ∀ {y:ℝ}, 0 < y → f (y + 1) = f y + log y) + (hn : n ≠ 0) (hx : 0 < x) (hx' : x ≤ 1) : + f (n + x) ≤ f n + x * log n := +begin + have hn': 0 < (n:ℝ) := nat.cast_pos.mpr (nat.pos_of_ne_zero hn), + have : f n + x * log n = (1 - x) * f n + x * f (n + 1), + { rw [hf_feq hn'], ring, }, + rw [this, (by ring : (n:ℝ) + x = (1 - x) * n + x * (n + 1))], + simpa only [smul_eq_mul] using hf_conv.2 hn' (by linarith : 0 < (n + 1 : ℝ)) + (by linarith : 0 ≤ 1 - x) hx.le (by linarith), +end + +/-- Linear lower bound for `f (x + n)` on unit interval -/ +lemma f_add_nat_ge + (hf_conv : convex_on ℝ (Ioi 0) f) (hf_feq : ∀ {y:ℝ}, 0 < y → f (y + 1) = f y + log y) + (hn : 2 ≤ n) (hx : 0 < x) : + f n + x * log (n - 1) ≤ f (n + x) := +begin + have npos : 0 < (n:ℝ) - 1, + { rw [←nat.cast_one, sub_pos, nat.cast_lt], linarith, }, + have c := (convex_on_iff_slope_mono_adjacent.mp $ hf_conv).2 + npos (by linarith : 0 < (n:ℝ) + x) (by linarith : (n:ℝ) - 1 < (n:ℝ)) (by linarith), + rw [add_sub_cancel', sub_sub_cancel, div_one] at c, + have : f (↑n - 1) = f n - log (↑n - 1), + { nth_rewrite_rhs 0 (by ring : (n:ℝ) = (↑n - 1) + 1), + rw [hf_feq npos, add_sub_cancel] }, + rwa [this, le_div_iff hx, sub_sub_cancel, le_sub_iff_add_le, mul_comm _ x, add_comm] at c, +end + +lemma log_gamma_seq_add_one (x : ℝ) (n : ℕ) : + log_gamma_seq (x + 1) n = log_gamma_seq x (n + 1) + log x - (x + 1) * (log (n + 1) - log n) := +begin + dsimp only [nat.factorial_succ, log_gamma_seq], + conv_rhs { rw [finset.sum_range_succ', nat.cast_zero, add_zero], }, + rw [nat.cast_mul, log_mul], rotate, + { rw nat.cast_ne_zero, exact nat.succ_ne_zero n }, + { rw nat.cast_ne_zero, exact nat.factorial_ne_zero n, }, + have : ∑ (m : ℕ) in finset.range (n + 1), log (x + 1 + ↑m) = + ∑ (k : ℕ) in finset.range (n + 1), log (x + ↑(k + 1)), + { refine finset.sum_congr (by refl) (λ m hm, _), + congr' 1, + push_cast, + abel }, + rw [←this, nat.cast_add_one n], + ring, +end + +lemma le_log_gamma_seq + (hf_conv : convex_on ℝ (Ioi 0) f) (hf_feq : ∀ {y:ℝ}, 0 < y → f (y + 1) = f y + log y) + (hx : 0 < x) (hx' : x ≤ 1) (n : ℕ) : + f x ≤ f 1 + x * log (n + 1) - x * log n + log_gamma_seq x n := +begin + rw [log_gamma_seq, ←add_sub_assoc, le_sub_iff_add_le, ←f_add_nat_eq @hf_feq hx, add_comm x], + refine (f_add_nat_le hf_conv @hf_feq (nat.add_one_ne_zero n) hx hx').trans (le_of_eq _), + rw [f_nat_eq @hf_feq (by linarith : n + 1 ≠ 0), nat.add_sub_cancel, nat.cast_add_one], + ring, +end + +lemma ge_log_gamma_seq + (hf_conv : convex_on ℝ (Ioi 0) f) (hf_feq : ∀ {y:ℝ}, 0 < y → f (y + 1) = f y + log y) + (hx : 0 < x) (hn : n ≠ 0) : + f 1 + log_gamma_seq x n ≤ f x := +begin + dsimp [log_gamma_seq], + rw [←add_sub_assoc, sub_le_iff_le_add, ←f_add_nat_eq @hf_feq hx, add_comm x _], + refine le_trans (le_of_eq _) (f_add_nat_ge hf_conv @hf_feq _ hx), + { rw [f_nat_eq @hf_feq, nat.add_sub_cancel, nat.cast_add_one, add_sub_cancel], + { ring }, + { exact nat.succ_ne_zero _} }, + { apply nat.succ_le_succ, + linarith [nat.pos_of_ne_zero hn] }, +end + +lemma tendsto_log_gamma_seq_of_le_one + (hf_conv : convex_on ℝ (Ioi 0) f) (hf_feq : ∀ {y:ℝ}, 0 < y → f (y + 1) = f y + log y) + (hx : 0 < x) (hx' : x ≤ 1) : + tendsto (log_gamma_seq x) at_top (𝓝 $ f x - f 1) := +begin + refine tendsto_of_tendsto_of_tendsto_of_le_of_le' _ tendsto_const_nhds _ _, + show ∀ᶠ (n : ℕ) in at_top, log_gamma_seq x n ≤ f x - f 1, + { refine eventually.mp (eventually_ne_at_top 0) (eventually_of_forall (λ n hn, _)), + exact le_sub_iff_add_le'.mpr (ge_log_gamma_seq hf_conv @hf_feq hx hn) }, + show ∀ᶠ (n : ℕ) in at_top, f x - f 1 - x * (log (n + 1) - log n) ≤ log_gamma_seq x n, + { refine eventually_of_forall (λ n, _), + rw [sub_le_iff_le_add', sub_le_iff_le_add'], + convert le_log_gamma_seq hf_conv @hf_feq hx hx' n using 1, + ring }, + { have : f x - f 1 = (f x - f 1) - x * 0 := by ring, + nth_rewrite 0 this, + exact tendsto.sub tendsto_const_nhds (tendsto_log_nat_add_one_sub_log.const_mul _), } +end + +lemma tendsto_log_gamma_seq + (hf_conv : convex_on ℝ (Ioi 0) f) (hf_feq : ∀ {y:ℝ}, 0 < y → f (y + 1) = f y + log y) + (hx : 0 < x) : + tendsto (log_gamma_seq x) at_top (𝓝 $ f x - f 1) := +begin + suffices : ∀ (m : ℕ), ↑m < x → x ≤ m + 1 → + tendsto (log_gamma_seq x) at_top (𝓝 $ f x - f 1), + { refine this (⌈x - 1⌉₊) _ _, + { rcases lt_or_le x 1, + { rwa [nat.ceil_eq_zero.mpr (by linarith : x - 1 ≤ 0), nat.cast_zero] }, + { convert nat.ceil_lt_add_one (by linarith : 0 ≤ x - 1), + abel } }, + { rw ←sub_le_iff_le_add, exact nat.le_ceil _}, }, + intro m, + induction m with m hm generalizing x, + { rw [nat.cast_zero, zero_add], + exact λ _ hx', tendsto_log_gamma_seq_of_le_one hf_conv @hf_feq hx hx' }, + { intros hy hy', + rw [nat.cast_succ, ←sub_le_iff_le_add] at hy', + rw [nat.cast_succ, ←lt_sub_iff_add_lt] at hy, + specialize hm ((nat.cast_nonneg _).trans_lt hy) hy hy', + -- now massage gauss_product n (x - 1) into gauss_product (n - 1) x + have : ∀ᶠ (n:ℕ) in at_top, log_gamma_seq (x - 1) n = log_gamma_seq x (n - 1) + + x * (log (↑(n - 1) + 1) - log ↑(n - 1)) - log (x - 1), + { refine eventually.mp (eventually_ge_at_top 1) (eventually_of_forall (λ n hn, _)), + have := log_gamma_seq_add_one (x - 1) (n - 1), + rw [sub_add_cancel, nat.sub_add_cancel hn] at this, + rw this, + ring }, + replace hm := ((tendsto.congr' this hm).add + (tendsto_const_nhds : tendsto (λ _, log (x - 1)) _ _)).comp (tendsto_add_at_top_nat 1), + have : + (λ (x_1 : ℕ), (λ (n : ℕ), log_gamma_seq x (n - 1) + + x * (log (↑(n - 1) + 1) - log ↑(n - 1)) - log (x - 1)) x_1 + + (λ (b : ℕ), log (x - 1)) x_1) ∘ (λ (a : ℕ), a + 1) = + λ n, log_gamma_seq x n + x * (log (↑n + 1) - log ↑n), + { ext1 n, + dsimp only [function.comp_app], + rw [sub_add_cancel, nat.add_sub_cancel] }, + rw this at hm, + convert hm.sub (tendsto_log_nat_add_one_sub_log.const_mul x) using 2, + { ext1 n, ring }, + { have := hf_feq ((nat.cast_nonneg m).trans_lt hy), + rw sub_add_cancel at this, + rw this, + ring } }, +end + +lemma tendsto_log_Gamma {x : ℝ} (hx : 0 < x) : + tendsto (log_gamma_seq x) at_top (𝓝 $ log (Gamma x)) := +begin + have : log (Gamma x) = (log ∘ Gamma) x - (log ∘ Gamma) 1, + { simp_rw [function.comp_app, Gamma_one, log_one, sub_zero] }, + rw this, + refine bohr_mollerup.tendsto_log_gamma_seq convex_on_log_Gamma (λ y hy, _) hx, + rw [function.comp_app, Gamma_add_one hy.ne', log_mul hy.ne' (Gamma_pos_of_pos hy).ne', add_comm], +end + +end bohr_mollerup -- (namespace) + +/-- The **Bohr-Mollerup theorem**: the Gamma function is the *unique* log-convex, positive-valued +function on the positive reals which satisfies `f 1 = 1` and `f (x + 1) = x * f x` for all `x`. -/ +lemma eq_Gamma_of_log_convex {f : ℝ → ℝ} + (hf_conv : convex_on ℝ (Ioi 0) (log ∘ f)) + (hf_feq : ∀ {y:ℝ}, 0 < y → f (y + 1) = y * f y) + (hf_pos : ∀ {y:ℝ}, 0 < y → 0 < f y) + (hf_one : f 1 = 1) : + eq_on f Gamma (Ioi (0:ℝ)) := +begin + suffices : eq_on (log ∘ f) (log ∘ Gamma) (Ioi (0:ℝ)), + from λ x hx, log_inj_on_pos (hf_pos hx) (Gamma_pos_of_pos hx) (this hx), + intros x hx, + have e1 := bohr_mollerup.tendsto_log_gamma_seq hf_conv _ hx, + { rw [function.comp_app log f 1, hf_one, log_one, sub_zero] at e1, + exact tendsto_nhds_unique e1 (bohr_mollerup.tendsto_log_Gamma hx) }, + { intros y hy, + rw [function.comp_app, hf_feq hy, log_mul hy.ne' (hf_pos hy).ne'], + ring } +end + +end bohr_mollerup -- (section) + +section strict_mono + +lemma Gamma_two : Gamma 2 = 1 := by simpa using Gamma_nat_eq_factorial 1 + +lemma Gamma_three_div_two_lt_one : Gamma (3 / 2) < 1 := +begin + -- This can also be proved using the closed-form evaluation of `Gamma (1 / 2)` in + -- `analysis.special_functions.gaussian`, but we give a self-contained proof using log-convexity + -- to avoid unnecessary imports. + have A : (0:ℝ) < 3/2, by norm_num, + have := bohr_mollerup.f_add_nat_le convex_on_log_Gamma (λ y hy, _) two_ne_zero one_half_pos + (by norm_num : 1/2 ≤ (1:ℝ)), + swap, { rw [function.comp_app, Gamma_add_one hy.ne', log_mul hy.ne' (Gamma_pos_of_pos hy).ne', + add_comm] }, + rw [function.comp_app, function.comp_app, nat.cast_two, Gamma_two, log_one, zero_add, + (by norm_num : (2:ℝ) + 1/2 = 3/2 + 1), Gamma_add_one A.ne', + log_mul A.ne' (Gamma_pos_of_pos A).ne', ←le_sub_iff_add_le', + log_le_iff_le_exp (Gamma_pos_of_pos A)] at this, + refine this.trans_lt (exp_lt_one_iff.mpr _), + rw [mul_comm, ←mul_div_assoc, div_sub' _ _ (2:ℝ) two_ne_zero], + refine div_neg_of_neg_of_pos _ two_pos, + rw [sub_neg, mul_one, ←nat.cast_two, ←log_pow, ←exp_lt_exp, nat.cast_two, exp_log two_pos, + exp_log]; + norm_num, +end + +lemma Gamma_strict_mono_on_Ici : strict_mono_on Gamma (Ici 2) := +begin + convert convex_on_Gamma.strict_mono_of_lt (by norm_num : (0:ℝ) < 3/2) + (by norm_num : (3/2 : ℝ) < 2) (Gamma_two.symm ▸ Gamma_three_div_two_lt_one), + symmetry, + rw inter_eq_right_iff_subset, + exact λ x hx, two_pos.trans_le hx, +end + +end strict_mono + +end real diff --git a/src/analysis/special_functions/gaussian.lean b/src/analysis/special_functions/gaussian.lean index 5825b1f8a1398..b21f166f048e6 100644 --- a/src/analysis/special_functions/gaussian.lean +++ b/src/analysis/special_functions/gaussian.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ -import analysis.special_functions.gamma +import analysis.special_functions.gamma.basic import analysis.special_functions.polar_coord import analysis.convex.complex import analysis.complex.cauchy_integral From 7ad820c4997738e2f542f8a20f32911f52020e26 Mon Sep 17 00:00:00 2001 From: Michael Stoll Date: Sun, 14 May 2023 15:14:09 +0000 Subject: [PATCH 1006/1291] doc(number_theory/pell): fix links in module docstring (#19010) This PR fixes two typos in the module docstring of `number_theory.pell` (a missing and a mis-spelt namespace part), which prevented doc-gen from linking to the relevant statements. --- src/number_theory/pell.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/number_theory/pell.lean b/src/number_theory/pell.lean index b8eae0615add6..53fde0c78ae84 100644 --- a/src/number_theory/pell.lean +++ b/src/number_theory/pell.lean @@ -28,13 +28,13 @@ We then prove the following **Theorem.** Let $d$ be a positive integer that is not a square. Then the equation $x^2 - d y^2 = 1$ has a nontrivial (i.e., with $y \ne 0$) solution in integers. -See `pell.exists_of_not_is_square` and `pell.exists_nontrivial_of_not_is_square`. +See `pell.exists_of_not_is_square` and `pell.solution₁.exists_nontrivial_of_not_is_square`. We then define the *fundamental solution* to be the solution with smallest $x$ among all solutions satisfying $x > 1$ and $y > 0$. We show that every solution is a power (in the sense of the group structure mentioned above) of the fundamental solution up to a (common) sign, -see `pell.fundamental.eq_zpow_or_neg_zpow`, and that a (positive) solution has this property +see `pell.is_fundamental.eq_zpow_or_neg_zpow`, and that a (positive) solution has this property if and only if it is fundamental, see `pell.pos_generator_iff_fundamental`. ## References From 2f5b500a507264de86d666a5f87ddb976e2d8de4 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sun, 14 May 2023 22:21:24 +0000 Subject: [PATCH 1007/1291] fix(data/mv_polynomial): add missing `decidable_eq` arguments to lemmas (#18848) This does not change the type of any definitions; the effect of this PR is to make the *statement* of the lemmas syntactically more general. To ensure this catches them all, this removes `open_locale classical` from the beginning of every file in `data/mv_polynomial` and `ring_theory/mv_polynomial`. For definitions which bake in a `classical.dec_eq` assumption, this adds a lemma proven by `convert rfl` that unfolds them to a version with an arbitrary `decidable_eq` instance, following a pattern established elsewhere. Unlike previous refactors of this style this doesn't seemed to have helped any downstream proofs much. --- src/data/mv_polynomial/basic.lean | 27 +++-- src/data/mv_polynomial/comm_ring.lean | 24 ++-- src/data/mv_polynomial/equiv.lean | 2 +- src/data/mv_polynomial/monad.lean | 6 +- src/data/mv_polynomial/pderiv.lean | 21 ++-- src/data/mv_polynomial/rename.lean | 8 +- src/data/mv_polynomial/supported.lean | 2 +- src/data/mv_polynomial/variables.lean | 110 ++++++++++++------ src/ring_theory/mv_polynomial/basic.lean | 6 +- .../mv_polynomial/homogeneous.lean | 4 +- src/ring_theory/mv_polynomial/symmetric.lean | 2 +- .../mv_polynomial/weighted_homogeneous.lean | 17 ++- src/ring_theory/witt_vector/mul_coeff.lean | 14 +-- 13 files changed, 153 insertions(+), 90 deletions(-) diff --git a/src/data/mv_polynomial/basic.lean b/src/data/mv_polynomial/basic.lean index 269d6e8156387..40641f917b012 100644 --- a/src/data/mv_polynomial/basic.lean +++ b/src/data/mv_polynomial/basic.lean @@ -72,7 +72,7 @@ polynomial, multivariate polynomial, multivariable polynomial noncomputable theory -open_locale classical big_operators +open_locale big_operators open set function finsupp add_monoid_algebra open_locale big_operators @@ -417,14 +417,14 @@ by convert rfl lemma support_monomial_subset : (monomial s a).support ⊆ {s} := support_single_subset -lemma support_add : (p + q).support ⊆ p.support ∪ q.support := finsupp.support_add +lemma support_add [decidable_eq σ] : (p + q).support ⊆ p.support ∪ q.support := finsupp.support_add lemma support_X [nontrivial R] : (X n : mv_polynomial σ R).support = {single n 1} := -by rw [X, support_monomial, if_neg]; exact one_ne_zero +by classical; rw [X, support_monomial, if_neg]; exact one_ne_zero lemma support_X_pow [nontrivial R] (s : σ) (n : ℕ) : (X s ^ n : mv_polynomial σ R).support = {finsupp.single s n} := -by rw [X_pow_eq_monomial, support_monomial, if_neg (one_ne_zero' R)] +by classical; rw [X_pow_eq_monomial, support_monomial, if_neg (one_ne_zero' R)] @[simp] lemma support_zero : (0 : mv_polynomial σ R).support = ∅ := rfl @@ -432,7 +432,7 @@ lemma support_smul {S₁ : Type*} [smul_zero_class S₁ R] {a : S₁} {f : mv_po (a • f).support ⊆ f.support := finsupp.support_smul -lemma support_sum {α : Type*} {s : finset α} {f : α → mv_polynomial σ R} : +lemma support_sum {α : Type*} [decidable_eq σ] {s : finset α} {f : α → mv_polynomial σ R} : (∑ x in s, f x).support ⊆ s.bUnion (λ x, (f x).support) := finsupp.support_finset_sum end support @@ -455,7 +455,7 @@ lemma sum_def {A} [add_comm_monoid A] {p : mv_polynomial σ R} {b : (σ →₀ p.sum b = ∑ m in p.support, b m (p.coeff m) := by simp [support, finsupp.sum, coeff] -lemma support_mul (p q : mv_polynomial σ R) : +lemma support_mul [decidable_eq σ] (p q : mv_polynomial σ R) : (p * q).support ⊆ p.support.bUnion (λ a, q.support.bUnion $ λ b, {a + b}) := by convert add_monoid_algebra.support_mul p q; ext; convert iff.rfl @@ -525,11 +525,12 @@ by rw [← coeff_X_pow, pow_one] @[simp] lemma coeff_X (i : σ) : coeff (single i 1) (X i : mv_polynomial σ R) = 1 := -by rw [coeff_X', if_pos rfl] +by classical; rw [coeff_X', if_pos rfl] @[simp] lemma coeff_C_mul (m) (a : R) (p : mv_polynomial σ R) : coeff m (C a * p) = a * coeff m p := begin + classical, rw [mul_def, sum_C], { simp [sum_def, coeff_sum] {contextual := tt} }, simp @@ -589,6 +590,7 @@ end lemma coeff_mul_monomial' (m) (s : σ →₀ ℕ) (r : R) (p : mv_polynomial σ R) : coeff m (p * monomial s r) = if s ≤ m then coeff (m - s) p * r else 0 := begin + classical, obtain rfl | hr := eq_or_ne r 0, { simp only [monomial_zero, coeff_zero, mul_zero, if_t_t], }, haveI : nontrivial R := nontrivial_of_ne _ _ hr, @@ -742,9 +744,12 @@ finsupp.sum_zero_index section @[simp] lemma eval₂_add : (p + q).eval₂ f g = p.eval₂ f g + q.eval₂ f g := -finsupp.sum_add_index - (by simp [f.map_zero]) - (by simp [add_mul, f.map_add]) +begin + classical, + exact finsupp.sum_add_index + (by simp [f.map_zero]) + (by simp [add_mul, f.map_add]) +end @[simp] lemma eval₂_monomial : (monomial s a).eval₂ f g = f a * s.prod (λn e, g n ^ e) := finsupp.sum_single_index (by simp [f.map_zero]) @@ -761,6 +766,7 @@ by simp [eval₂_monomial, f.map_one, X, prod_single_index, pow_one] lemma eval₂_mul_monomial : ∀{s a}, (p * monomial s a).eval₂ f g = p.eval₂ f g * f a * s.prod (λn e, g n ^ e) := begin + classical, apply mv_polynomial.induction_on p, { assume a' s a, simp [C_mul_monomial, eval₂_monomial, f.map_mul] }, @@ -994,6 +1000,7 @@ end lemma coeff_map (p : mv_polynomial σ R) : ∀ (m : σ →₀ ℕ), coeff m (map f p) = f (coeff m p) := begin + classical, apply mv_polynomial.induction_on p; clear p, { intros r m, rw [map_C], simp only [coeff_C], split_ifs, {refl}, rw f.map_zero }, { intros p q hp hq m, simp only [hp, hq, (map f).map_add, coeff_add], rw f.map_add }, diff --git a/src/data/mv_polynomial/comm_ring.lean b/src/data/mv_polynomial/comm_ring.lean index a7ccd331d6314..3d8107b69fca1 100644 --- a/src/data/mv_polynomial/comm_ring.lean +++ b/src/data/mv_polynomial/comm_ring.lean @@ -38,8 +38,6 @@ This will give rise to a monomial in `mv_polynomial σ R` which mathematicians m noncomputable theory -open_locale classical big_operators - open set function finsupp add_monoid_algebra open_locale big_operators @@ -70,7 +68,8 @@ variables (σ a a') @[simp] lemma support_neg : (- p).support = p.support := finsupp.support_neg p -lemma support_sub (p q : mv_polynomial σ R) : (p - q).support ⊆ p.support ∪ q.support := +lemma support_sub [decidable_eq σ] (p q : mv_polynomial σ R) : + (p - q).support ⊆ p.support ∪ q.support := finsupp.support_sub variables {σ} (p) @@ -80,9 +79,9 @@ section degrees lemma degrees_neg (p : mv_polynomial σ R) : (- p).degrees = p.degrees := by rw [degrees, support_neg]; refl -lemma degrees_sub (p q : mv_polynomial σ R) : +lemma degrees_sub [decidable_eq σ] (p q : mv_polynomial σ R) : (p - q).degrees ≤ p.degrees ⊔ q.degrees := -by simpa only [sub_eq_add_neg] using le_trans (degrees_add p (-q)) (by rw degrees_neg) +by classical; simpa only [sub_eq_add_neg, degrees_neg] using degrees_add p (-q) end degrees @@ -93,13 +92,14 @@ variables (p q) @[simp] lemma vars_neg : (-p).vars = p.vars := by simp [vars, degrees_neg] -lemma vars_sub_subset : (p - q).vars ⊆ p.vars ∪ q.vars := +lemma vars_sub_subset [decidable_eq σ] : (p - q).vars ⊆ p.vars ∪ q.vars := by convert vars_add_subset p (-q) using 2; simp [sub_eq_add_neg] variables {p q} @[simp] -lemma vars_sub_of_disjoint (hpq : disjoint p.vars q.vars) : (p - q).vars = p.vars ∪ q.vars := +lemma vars_sub_of_disjoint [decidable_eq σ] (hpq : disjoint p.vars q.vars) : + (p - q).vars = p.vars ∪ q.vars := begin rw ←vars_neg q at hpq, convert vars_add_of_disjoint hpq using 2; @@ -148,6 +148,7 @@ lemma degree_of_sub_lt {x : σ} {f g : mv_polynomial σ R} {k : ℕ} (h : 0 < k) (hg : ∀ (m : σ →₀ ℕ), m ∈ g.support → (k ≤ m x) → coeff m f = coeff m g) : degree_of x (f - g) < k := begin + classical, rw degree_of_lt_iff h, intros m hm, by_contra hc, @@ -169,9 +170,12 @@ by simp only [total_degree, support_neg] lemma total_degree_sub (a b : mv_polynomial σ R) : (a - b).total_degree ≤ max a.total_degree b.total_degree := -calc (a - b).total_degree = (a + -b).total_degree : by rw sub_eq_add_neg - ... ≤ max a.total_degree (-b).total_degree : total_degree_add a (-b) - ... = max a.total_degree b.total_degree : by rw total_degree_neg +begin + classical, + calc (a - b).total_degree = (a + -b).total_degree : by rw sub_eq_add_neg + ... ≤ max a.total_degree (-b).total_degree : total_degree_add a (-b) + ... = max a.total_degree b.total_degree : by rw total_degree_neg +end end total_degree diff --git a/src/data/mv_polynomial/equiv.lean b/src/data/mv_polynomial/equiv.lean index f4dadd5a7c302..6fc617c5ad0fe 100644 --- a/src/data/mv_polynomial/equiv.lean +++ b/src/data/mv_polynomial/equiv.lean @@ -46,7 +46,7 @@ equivalence, isomorphism, morphism, ring hom, hom noncomputable theory -open_locale classical big_operators polynomial +open_locale big_operators polynomial open set function finsupp add_monoid_algebra diff --git a/src/data/mv_polynomial/monad.lean b/src/data/mv_polynomial/monad.lean index 106cefe8a1c1d..ca12e69353d15 100644 --- a/src/data/mv_polynomial/monad.lean +++ b/src/data/mv_polynomial/monad.lean @@ -288,9 +288,8 @@ lemma bind₂_monomial_one (f : R →+* mv_polynomial σ S) (d : σ →₀ ℕ) by rw [bind₂_monomial, f.map_one, one_mul] section -open_locale classical -lemma vars_bind₁ (f : σ → mv_polynomial τ R) (φ : mv_polynomial σ R) : +lemma vars_bind₁ [decidable_eq τ] (f : σ → mv_polynomial τ R) (φ : mv_polynomial σ R) : (bind₁ f φ).vars ⊆ φ.vars.bUnion (λ i, (f i).vars) := begin calc (bind₁ f φ).vars @@ -323,7 +322,8 @@ end lemma mem_vars_bind₁ (f : σ → mv_polynomial τ R) (φ : mv_polynomial σ R) {j : τ} (h : j ∈ (bind₁ f φ).vars) : ∃ (i : σ), i ∈ φ.vars ∧ j ∈ (f i).vars := -by simpa only [exists_prop, finset.mem_bUnion, mem_support_iff, ne.def] using vars_bind₁ f φ h +by classical; simpa only [exists_prop, finset.mem_bUnion, mem_support_iff, ne.def] + using vars_bind₁ f φ h instance monad : monad (λ σ, mv_polynomial σ R) := { map := λ α β f p, rename f p, diff --git a/src/data/mv_polynomial/pderiv.lean b/src/data/mv_polynomial/pderiv.lean index 4cd9c1e640697..1e95dac506240 100644 --- a/src/data/mv_polynomial/pderiv.lean +++ b/src/data/mv_polynomial/pderiv.lean @@ -47,7 +47,7 @@ namespace mv_polynomial open set function finsupp add_monoid_algebra -open_locale classical big_operators +open_locale big_operators variables {R : Type u} {σ : Type v} {a a' a₁ a₂ : R} {s : σ →₀ ℕ} @@ -57,12 +57,16 @@ variables {R} [comm_semiring R] /-- `pderiv i p` is the partial derivative of `p` with respect to `i` -/ def pderiv (i : σ) : derivation R (mv_polynomial σ R) (mv_polynomial σ R) := -mk_derivation R $ pi.single i 1 +by letI := classical.dec_eq σ; exact (mk_derivation R $ pi.single i 1) + +lemma pderiv_def [decidable_eq σ] (i : σ) : pderiv i = mk_derivation R (pi.single i 1) := +by convert rfl @[simp] lemma pderiv_monomial {i : σ} : pderiv i (monomial s a) = monomial (s - single i 1) (a * (s i)) := begin - simp only [pderiv, mk_derivation_monomial, finsupp.smul_sum, smul_eq_mul, + classical, + simp only [pderiv_def, mk_derivation_monomial, finsupp.smul_sum, smul_eq_mul, ← smul_mul_assoc, ← (monomial _).map_smul], refine (finset.sum_eq_single i (λ j hj hne, _) (λ hi, _)).trans _, { simp [pi.single_eq_of_ne hne] }, @@ -74,14 +78,15 @@ lemma pderiv_C {i : σ} : pderiv i (C a) = 0 := derivation_C _ _ lemma pderiv_one {i : σ} : pderiv i (1 : mv_polynomial σ R) = 0 := pderiv_C -@[simp] lemma pderiv_X [d : decidable_eq σ] (i j : σ) : - pderiv i (X j : mv_polynomial σ R) = @pi.single σ _ d _ i 1 j := -(mk_derivation_X _ _ _).trans (by congr) +@[simp] lemma pderiv_X [decidable_eq σ] (i j : σ) : + pderiv i (X j : mv_polynomial σ R) = @pi.single _ _ _ _ i 1 j := +by rw [pderiv_def, mk_derivation_X] -@[simp] lemma pderiv_X_self (i : σ) : pderiv i (X i : mv_polynomial σ R) = 1 := by simp +@[simp] lemma pderiv_X_self (i : σ) : pderiv i (X i : mv_polynomial σ R) = 1 := +by classical; simp @[simp] lemma pderiv_X_of_ne {i j : σ} (h : j ≠ i) : pderiv i (X j : mv_polynomial σ R) = 0 := -by simp [h] +by classical; simp [h] lemma pderiv_eq_zero_of_not_mem_vars {i : σ} {f : mv_polynomial σ R} (h : i ∉ f.vars) : pderiv i f = 0 := diff --git a/src/data/mv_polynomial/rename.lean b/src/data/mv_polynomial/rename.lean index 858e5b2cf69d0..95a4074a22682 100644 --- a/src/data/mv_polynomial/rename.lean +++ b/src/data/mv_polynomial/rename.lean @@ -41,7 +41,7 @@ This will give rise to a monomial in `mv_polynomial σ R` which mathematicians m noncomputable theory -open_locale classical big_operators +open_locale big_operators open set function finsupp add_monoid_algebra open_locale big_operators @@ -185,6 +185,7 @@ end theorem exists_finset_rename (p : mv_polynomial σ R) : ∃ (s : finset σ) (q : mv_polynomial {x // x ∈ s} R), p = rename coe q := begin + classical, apply induction_on p, { intro r, exact ⟨∅, C r, by rw rename_C⟩ }, { rintro p q ⟨s, p, rfl⟩ ⟨t, q, rfl⟩, @@ -241,6 +242,7 @@ section coeff lemma coeff_rename_map_domain (f : σ → τ) (hf : injective f) (φ : mv_polynomial σ R) (d : σ →₀ ℕ) : (rename f φ).coeff (d.map_domain f) = φ.coeff d := begin + classical, apply induction_on' φ, { intros u r, rw [rename_monomial, coeff_monomial, coeff_monomial], @@ -252,6 +254,7 @@ lemma coeff_rename_eq_zero (f : σ → τ) (φ : mv_polynomial σ R) (d : τ → (h : ∀ u : σ →₀ ℕ, u.map_domain f = d → φ.coeff u = 0) : (rename f φ).coeff d = 0 := begin + classical, rw [rename_eq, ← not_mem_support_iff], intro H, replace H := map_domain_support H, @@ -280,7 +283,8 @@ end coeff section support -lemma support_rename_of_injective {p : mv_polynomial σ R} {f : σ → τ} (h : function.injective f) : +lemma support_rename_of_injective {p : mv_polynomial σ R} {f : σ → τ} [decidable_eq τ] + (h : function.injective f) : (rename f p).support = finset.image (map_domain f) p.support := begin rw rename_eq, diff --git a/src/data/mv_polynomial/supported.lean b/src/data/mv_polynomial/supported.lean index 41ce4104cbf5c..e1f5a2998ad54 100644 --- a/src/data/mv_polynomial/supported.lean +++ b/src/data/mv_polynomial/supported.lean @@ -38,7 +38,6 @@ algebra.adjoin R (X '' s) variables {σ R} -open_locale classical open algebra lemma supported_eq_range_rename (s : set σ) : @@ -67,6 +66,7 @@ variables {s t : set σ} lemma mem_supported : p ∈ (supported R s) ↔ ↑p.vars ⊆ s := begin + classical, rw [supported_eq_range_rename, alg_hom.mem_range], split, { rintros ⟨p, rfl⟩, diff --git a/src/data/mv_polynomial/variables.lean b/src/data/mv_polynomial/variables.lean index bcbcca71ae509..9fd1ca013972e 100644 --- a/src/data/mv_polynomial/variables.lean +++ b/src/data/mv_polynomial/variables.lean @@ -58,8 +58,6 @@ This will give rise to a monomial in `mv_polynomial σ R` which mathematicians m noncomputable theory -open_locale classical big_operators - open set function finsupp add_monoid_algebra open_locale big_operators @@ -83,11 +81,15 @@ The maximal degrees of each variable in a multi-variable polynomial, expressed a (For example, `degrees (x^2 * y + y^3)` would be `{x, x, y, y, y}`.) -/ def degrees (p : mv_polynomial σ R) : multiset σ := -p.support.sup (λs:σ →₀ ℕ, s.to_multiset) +by letI := classical.dec_eq σ; exact p.support.sup (λs:σ →₀ ℕ, s.to_multiset) + +lemma degrees_def [decidable_eq σ] (p : mv_polynomial σ R) : + p.degrees = p.support.sup (λs:σ →₀ ℕ, s.to_multiset) := by convert rfl lemma degrees_monomial (s : σ →₀ ℕ) (a : R) : degrees (monomial s a) ≤ s.to_multiset := -finset.sup_le $ assume t h, begin + classical, + refine (finset.sup_le $ assume t h, _), have := finsupp.support_single_subset h, rw [finset.mem_singleton] at this, rw this @@ -95,8 +97,11 @@ end lemma degrees_monomial_eq (s : σ →₀ ℕ) (a : R) (ha : a ≠ 0) : degrees (monomial s a) = s.to_multiset := -le_antisymm (degrees_monomial s a) $ finset.le_sup $ - by rw [support_monomial, if_neg ha, finset.mem_singleton] +begin + classical, + refine (le_antisymm (degrees_monomial s a) $ finset.le_sup $ _), + rw [support_monomial, if_neg ha, finset.mem_singleton] +end lemma degrees_C (a : R) : degrees (C a : mv_polynomial σ R) = 0 := multiset.le_zero.1 $ degrees_monomial _ _ @@ -112,8 +117,11 @@ by { rw ← C_0, exact degrees_C 0 } @[simp] lemma degrees_one : degrees (1 : mv_polynomial σ R) = 0 := degrees_C 1 -lemma degrees_add (p q : mv_polynomial σ R) : (p + q).degrees ≤ p.degrees ⊔ q.degrees := +lemma degrees_add [decidable_eq σ] (p q : mv_polynomial σ R) : + (p + q).degrees ≤ p.degrees ⊔ q.degrees := begin + classical, + simp_rw degrees_def, refine finset.sup_le (assume b hb, _), have := finsupp.support_add hb, rw finset.mem_union at this, cases this, @@ -121,9 +129,10 @@ begin { exact le_sup_of_le_right (finset.le_sup this) }, end -lemma degrees_sum {ι : Type*} (s : finset ι) (f : ι → mv_polynomial σ R) : +lemma degrees_sum {ι : Type*} [decidable_eq σ] (s : finset ι) (f : ι → mv_polynomial σ R) : (∑ i in s, f i).degrees ≤ s.sup (λi, (f i).degrees) := begin + classical, refine s.induction _ _, { simp only [finset.sum_empty, finset.sup_empty, degrees_zero], exact le_rfl }, { assume i s his ih, @@ -133,6 +142,7 @@ end lemma degrees_mul (p q : mv_polynomial σ R) : (p * q).degrees ≤ p.degrees + q.degrees := begin + classical, refine finset.sup_le (assume b hb, _), have := support_mul p q hb, simp only [finset.mem_bUnion, finset.mem_singleton] at this, @@ -144,6 +154,7 @@ end lemma degrees_prod {ι : Type*} (s : finset ι) (f : ι → mv_polynomial σ R) : (∏ i in s, f i).degrees ≤ ∑ i in s, (f i).degrees := begin + classical, refine s.induction _ _, { simp only [finset.prod_empty, finset.sum_empty, degrees_one] }, { assume i s his ih, @@ -165,6 +176,7 @@ by simp only [degrees, multiset.mem_sup, ← mem_support_iff, lemma le_degrees_add {p q : mv_polynomial σ R} (h : p.degrees.disjoint q.degrees) : p.degrees ≤ (p + q).degrees := begin + classical, apply finset.sup_le, intros d hd, rw multiset.disjoint_iff_ne at h, @@ -186,7 +198,7 @@ begin all_goals { rw mem_degrees, refine ⟨d, _, hj⟩, assumption } } end -lemma degrees_add_of_disjoint +lemma degrees_add_of_disjoint [decidable_eq σ] {p q : mv_polynomial σ R} (h : multiset.disjoint p.degrees q.degrees) : (p + q).degrees = p.degrees ∪ q.degrees := begin @@ -209,6 +221,7 @@ end lemma degrees_rename (f : σ → τ) (φ : mv_polynomial σ R) : (rename f φ).degrees ⊆ (φ.degrees.map f) := begin + classical, intros i, rw [mem_degrees, multiset.mem_map], rintro ⟨d, hd, hi⟩, @@ -230,6 +243,7 @@ by simp only [degrees, mv_polynomial.support_map_of_injective _ hf] lemma degrees_rename_of_injective {p : mv_polynomial σ R} {f : σ → τ} (h : function.injective f) : degrees (rename f p) = (degrees p).map f := begin + classical, simp only [degrees, multiset.map_finset_sup p.support finsupp.to_multiset f h, support_rename_of_injective h, finset.sup_image], refine finset.sup_congr rfl (λ x hx, _), @@ -243,16 +257,20 @@ section vars /-! ### `vars` -/ /-- `vars p` is the set of variables appearing in the polynomial `p` -/ -def vars (p : mv_polynomial σ R) : finset σ := p.degrees.to_finset +def vars (p : mv_polynomial σ R) : finset σ := +by letI := classical.dec_eq σ; exact p.degrees.to_finset + +lemma vars_def [decidable_eq σ] (p : mv_polynomial σ R) : p.vars = p.degrees.to_finset := +by convert rfl @[simp] lemma vars_0 : (0 : mv_polynomial σ R).vars = ∅ := -by rw [vars, degrees_zero, multiset.to_finset_zero] +by classical; rw [vars_def, degrees_zero, multiset.to_finset_zero] @[simp] lemma vars_monomial (h : r ≠ 0) : (monomial s r).vars = s.support := -by rw [vars, degrees_monomial_eq _ _ h, finsupp.to_finset_to_multiset] +by classical; rw [vars_def, degrees_monomial_eq _ _ h, finsupp.to_finset_to_multiset] @[simp] lemma vars_C : (C r : mv_polynomial σ R).vars = ∅ := -by rw [vars, degrees_C, multiset.to_finset_zero] +by classical; rw [vars_def, degrees_C, multiset.to_finset_zero] @[simp] lemma vars_X [nontrivial R] : (X n : mv_polynomial σ R).vars = {n} := by rw [X, vars_monomial (one_ne_zero' R), finsupp.support_single_ne_zero _ (one_ne_zero' ℕ)] @@ -266,10 +284,11 @@ lemma mem_support_not_mem_vars_zero {f : mv_polynomial σ R} {x : σ →₀ ℕ} (H : x ∈ f.support) {v : σ} (h : v ∉ vars f) : x v = 0 := begin - rw [vars, multiset.mem_to_finset] at h, + letI := classical.dec_eq σ, + rw [vars_def, multiset.mem_to_finset] at h, rw ← finsupp.not_mem_support_iff, contrapose! h, - unfold degrees, + rw degrees_def, rw (show f.support = insert x f.support, from eq.symm $ finset.insert_eq_of_mem H), rw finset.sup_insert, simp only [multiset.mem_union, multiset.sup_eq_union], @@ -277,7 +296,7 @@ begin rwa [←to_finset_to_multiset, multiset.mem_to_finset] at h, end -lemma vars_add_subset (p q : mv_polynomial σ R) : +lemma vars_add_subset [decidable_eq σ] (p q : mv_polynomial σ R) : (p + q).vars ⊆ p.vars ∪ q.vars := begin intros x hx, @@ -285,19 +304,19 @@ begin simpa using multiset.mem_of_le (degrees_add _ _) hx, end -lemma vars_add_of_disjoint (h : disjoint p.vars q.vars) : +lemma vars_add_of_disjoint [decidable_eq σ] (h : disjoint p.vars q.vars) : (p + q).vars = p.vars ∪ q.vars := begin apply finset.subset.antisymm (vars_add_subset p q), intros x hx, - simp only [vars, multiset.disjoint_to_finset] at h hx ⊢, + simp only [vars_def, multiset.disjoint_to_finset] at h hx ⊢, rw [degrees_add_of_disjoint h, multiset.to_finset_union], exact hx end section mul -lemma vars_mul (φ ψ : mv_polynomial σ R) : (φ * ψ).vars ⊆ φ.vars ∪ ψ.vars := +lemma vars_mul [decidable_eq σ] (φ ψ : mv_polynomial σ R) : (φ * ψ).vars ⊆ φ.vars ∪ ψ.vars := begin intro i, simp only [mem_vars, finset.mem_union], @@ -321,6 +340,8 @@ vars_C lemma vars_pow (φ : mv_polynomial σ R) (n : ℕ) : (φ ^ n).vars ⊆ φ.vars := begin + classical, + simp_rw vars_def, induction n with n ih, { simp }, { rw pow_succ, @@ -332,9 +353,10 @@ end The variables of the product of a family of polynomials are a subset of the union of the sets of variables of each polynomial. -/ -lemma vars_prod {ι : Type*} {s : finset ι} (f : ι → mv_polynomial σ R) : +lemma vars_prod {ι : Type*} [decidable_eq σ] {s : finset ι} (f : ι → mv_polynomial σ R) : (∏ i in s, f i).vars ⊆ s.bUnion (λ i, (f i).vars) := begin + classical, apply s.induction_on, { simp }, { intros a s hs hsub, @@ -364,9 +386,10 @@ section sum variables {ι : Type*} (t : finset ι) (φ : ι → mv_polynomial σ R) -lemma vars_sum_subset : +lemma vars_sum_subset [decidable_eq σ] : (∑ i in t, φ i).vars ⊆ finset.bUnion t (λ i, (φ i).vars) := begin + classical, apply t.induction_on, { simp }, { intros a s has hsum, @@ -376,9 +399,10 @@ begin assumption } end -lemma vars_sum_of_disjoint (h : pairwise $ disjoint on (λ i, (φ i).vars)) : +lemma vars_sum_of_disjoint [decidable_eq σ] (h : pairwise $ disjoint on (λ i, (φ i).vars)) : (∑ i in t, φ i).vars = finset.bUnion t (λ i, (φ i).vars) := begin + classical, apply t.induction_on, { simp }, { intros a s has hsum, @@ -413,7 +437,7 @@ lemma vars_monomial_single (i : σ) {e : ℕ} {r : R} (he : e ≠ 0) (hr : r ≠ (monomial (finsupp.single i e) r).vars = {i} := by rw [vars_monomial hr, finsupp.support_single_ne_zero _ he] -lemma vars_eq_support_bUnion_support : p.vars = p.support.bUnion finsupp.support := +lemma vars_eq_support_bUnion_support [decidable_eq σ] : p.vars = p.support.bUnion finsupp.support := by { ext i, rw [mem_vars, finset.mem_bUnion] } end map @@ -425,12 +449,18 @@ section degree_of /-! ### `degree_of` -/ /-- `degree_of n p` gives the highest power of X_n that appears in `p` -/ -def degree_of (n : σ) (p : mv_polynomial σ R) : ℕ := p.degrees.count n +def degree_of (n : σ) (p : mv_polynomial σ R) : ℕ := +by letI := classical.dec_eq σ; exact p.degrees.count n + +lemma degree_of_def [decidable_eq σ] (n : σ) (p : mv_polynomial σ R) : + p.degree_of n = p.degrees.count n := +by convert rfl lemma degree_of_eq_sup (n : σ) (f : mv_polynomial σ R) : degree_of n f = f.support.sup (λ m, m n) := begin - rw [degree_of, degrees, multiset.count_finset_sup], + classical, + rw [degree_of_def, degrees_def, multiset.count_finset_sup], congr, ext, simp, @@ -447,7 +477,7 @@ by simp only [degree_of, degrees_zero, multiset.count_zero] @[simp] lemma degree_of_C (a : R) (x : σ): degree_of x (C a : mv_polynomial σ R) = 0 := by simp [degree_of, degrees_C] -lemma degree_of_X (i j : σ) [nontrivial R] : +lemma degree_of_X [decidable_eq σ] (i j : σ) [nontrivial R] : degree_of i (X j : mv_polynomial σ R) = if i = j then 1 else 0 := begin by_cases c : i = j, @@ -458,7 +488,8 @@ end lemma degree_of_add_le (n : σ) (f g : mv_polynomial σ R) : degree_of n (f + g) ≤ max (degree_of n f) (degree_of n g) := begin - repeat {rw degree_of}, + classical, + repeat {rw degree_of_def}, apply (multiset.count_le_of_le n (degrees_add f g)).trans, dsimp, rw multiset.count_union, @@ -475,7 +506,8 @@ end lemma degree_of_mul_le (i : σ) (f g: mv_polynomial σ R) : degree_of i (f * g) ≤ degree_of i f + degree_of i g := begin - repeat {rw degree_of}, + classical, + repeat {rw degree_of_def}, convert multiset.count_le_of_le i (degrees_mul f g), rw multiset.count_add, end @@ -483,6 +515,7 @@ end lemma degree_of_mul_X_ne {i j : σ} (f : mv_polynomial σ R) (h : i ≠ j) : degree_of i (f * X j) = degree_of i f := begin + classical, repeat {rw degree_of_eq_sup i}, rw support_mul_X, simp only [finset.sup_map], @@ -496,7 +529,8 @@ end lemma degree_of_mul_X_eq (j : σ) (f : mv_polynomial σ R) : degree_of j (f * X j) ≤ degree_of j f + 1 := begin - repeat {rw degree_of}, + classical, + repeat {rw degree_of_def}, apply (multiset.count_le_of_le j (degrees_mul f (X j))).trans, simp only [multiset.count_add, add_le_add_iff_left], convert multiset.count_le_of_le j (degrees_X' j), @@ -505,8 +539,8 @@ end lemma degree_of_rename_of_injective {p : mv_polynomial σ R} {f : σ → τ} (h : function.injective f) (i : σ) : degree_of (f i) (rename f p) = degree_of i p := -by simp only [degree_of, degrees_rename_of_injective h, - multiset.count_map_eq_count' f (p.degrees) h] +by classical; simp only [degree_of_def, degrees_rename_of_injective h, + multiset.count_map_eq_count' f (p.degrees) h] end degree_of @@ -528,6 +562,7 @@ end lemma total_degree_le_degrees_card (p : mv_polynomial σ R) : p.total_degree ≤ p.degrees.card := begin + classical, rw [total_degree_eq], exact finset.sup_le (assume s hs, multiset.card_le_of_le $ finset.le_sup hs) end @@ -557,8 +592,9 @@ end lemma total_degree_add (a b : mv_polynomial σ R) : (a + b).total_degree ≤ max a.total_degree b.total_degree := finset.sup_le $ assume n hn, - have _ := finsupp.support_add hn, begin + classical, + have := finsupp.support_add hn, rw finset.mem_union at this, cases this, { exact le_max_of_le_left (finset.le_sup this) }, @@ -595,8 +631,9 @@ by rw [add_comm, total_degree_add_eq_left_of_total_degree_lt h] lemma total_degree_mul (a b : mv_polynomial σ R) : (a * b).total_degree ≤ a.total_degree + b.total_degree := finset.sup_le $ assume n hn, - have _ := add_monoid_algebra.support_mul a b hn, begin + classical, + have := add_monoid_algebra.support_mul a b hn, simp only [finset.mem_bUnion, finset.mem_singleton] at this, rcases this with ⟨a₁, h₁, a₂, h₂, rfl⟩, rw [finsupp.sum_add_index'], @@ -623,7 +660,7 @@ end @[simp] lemma total_degree_monomial (s : σ →₀ ℕ) {c : R} (hc : c ≠ 0) : (monomial s c : mv_polynomial σ R).total_degree = s.sum (λ _ e, e) := -by simp [total_degree, support_monomial, if_neg hc] +by classical; simp [total_degree, support_monomial, if_neg hc] @[simp] lemma total_degree_X_pow [nontrivial R] (s : σ) (n : ℕ) : (X s ^ n : mv_polynomial σ R).total_degree = n := @@ -693,6 +730,7 @@ finset.sup_le $ assume b, begin assume h, rw rename_eq at h, + classical, have h' := finsupp.map_domain_support h, rw finset.mem_image at h', rcases h' with ⟨s, hs, rfl⟩, @@ -788,7 +826,7 @@ lemma exists_rename_eq_of_vars_subset_range { refl } end⟩ -lemma vars_rename (f : σ → τ) (φ : mv_polynomial σ R) : +lemma vars_rename [decidable_eq τ] (f : σ → τ) (φ : mv_polynomial σ R) : (rename f φ).vars ⊆ (φ.vars.image f) := begin intros i hi, @@ -798,7 +836,7 @@ end lemma mem_vars_rename (f : σ → τ) (φ : mv_polynomial σ R) {j : τ} (h : j ∈ (rename f φ).vars) : ∃ (i : σ), i ∈ φ.vars ∧ f i = j := -by simpa only [exists_prop, finset.mem_image] using vars_rename f φ h +by classical; simpa only [exists_prop, finset.mem_image] using vars_rename f φ h end eval_vars diff --git a/src/ring_theory/mv_polynomial/basic.lean b/src/ring_theory/mv_polynomial/basic.lean index 597e65023f10f..6d68cb3780924 100644 --- a/src/ring_theory/mv_polynomial/basic.lean +++ b/src/ring_theory/mv_polynomial/basic.lean @@ -39,8 +39,6 @@ Generalise to noncommutative (semi)rings noncomputable theory -open_locale classical - open set linear_map submodule open_locale big_operators polynomial @@ -99,10 +97,10 @@ begin refl end -lemma mem_restrict_degree_iff_sup (p : mv_polynomial σ R) (n : ℕ) : +lemma mem_restrict_degree_iff_sup [decidable_eq σ] (p : mv_polynomial σ R) (n : ℕ) : p ∈ restrict_degree σ R n ↔ ∀i, p.degrees.count i ≤ n := begin - simp only [mem_restrict_degree, degrees, multiset.count_finset_sup, finsupp.count_to_multiset, + simp only [mem_restrict_degree, degrees_def, multiset.count_finset_sup, finsupp.count_to_multiset, finset.sup_le_iff], exact ⟨assume h n s hs, h s hs n, assume h s hs n, h n s hs⟩ end diff --git a/src/ring_theory/mv_polynomial/homogeneous.lean b/src/ring_theory/mv_polynomial/homogeneous.lean index b19dda3233d45..40f0e3987b7ce 100644 --- a/src/ring_theory/mv_polynomial/homogeneous.lean +++ b/src/ring_theory/mv_polynomial/homogeneous.lean @@ -229,7 +229,6 @@ end is_homogeneous section noncomputable theory -open_locale classical open finset /-- `homogeneous_component n φ` is the part of `φ` that is homogeneous of degree `n`. @@ -265,7 +264,8 @@ lemma homogeneous_component_zero : homogeneous_component 0 φ = C (coeff 0 φ) : begin ext1 d, rcases em (d = 0) with (rfl|hd), - { simp only [coeff_homogeneous_component, sum_eq_zero_iff, finsupp.zero_apply, if_true, coeff_C, + { classical, + simp only [coeff_homogeneous_component, sum_eq_zero_iff, finsupp.zero_apply, if_true, coeff_C, eq_self_iff_true, forall_true_iff] }, { rw [coeff_homogeneous_component, if_neg, coeff_C, if_neg (ne.symm hd)], simp only [finsupp.ext_iff, finsupp.zero_apply] at hd, diff --git a/src/ring_theory/mv_polynomial/symmetric.lean b/src/ring_theory/mv_polynomial/symmetric.lean index d0ab30e3b7176..9eb7b27d26b8f 100644 --- a/src/ring_theory/mv_polynomial/symmetric.lean +++ b/src/ring_theory/mv_polynomial/symmetric.lean @@ -215,7 +215,7 @@ begin classical, have : (finsupp.to_multiset ∘ λ (t : finset σ), ∑ (i : σ) in t, finsupp.single i 1) = finset.val, { funext, simp [finsupp.to_multiset_sum_single] }, - rw [degrees, support_esymm, sup_image, this, ←comp_sup_eq_sup_comp], + rw [degrees_def, support_esymm, sup_image, this, ←comp_sup_eq_sup_comp], { obtain ⟨k, rfl⟩ := nat.exists_eq_succ_of_ne_zero hpos.ne', simpa using powerset_len_sup _ _ (nat.lt_of_succ_le hn) }, { intros, diff --git a/src/ring_theory/mv_polynomial/weighted_homogeneous.lean b/src/ring_theory/mv_polynomial/weighted_homogeneous.lean index 5e16ce5bad6d8..6def165740d0f 100644 --- a/src/ring_theory/mv_polynomial/weighted_homogeneous.lean +++ b/src/ring_theory/mv_polynomial/weighted_homogeneous.lean @@ -43,7 +43,7 @@ components. noncomputable theory -open_locale classical big_operators +open_locale big_operators open set function finset finsupp add_monoid_algebra @@ -185,6 +185,7 @@ end lemma is_weighted_homogeneous_monomial (w : σ → M) (d : σ →₀ ℕ) (r : R) {m : M} (hm : weighted_degree' w d = m) : is_weighted_homogeneous w (monomial d r) m := begin + classical, intros c hc, rw coeff_monomial at hc, split_ifs at hc with h, @@ -270,6 +271,7 @@ lemma prod {ι : Type*} (s : finset ι) (φ : ι → mv_polynomial σ R) (n : ι (∀ i ∈ s, is_weighted_homogeneous w (φ i) (n i)) → is_weighted_homogeneous w (∏ i in s, φ i) (∑ i in s, n i) := begin + classical, apply finset.induction_on s, { intro, simp only [is_weighted_homogeneous_one, finset.sum_empty, finset.prod_empty] }, { intros i s his IH h, @@ -316,12 +318,12 @@ section weighted_homogeneous_component variables {w : σ → M} (n : M) (φ ψ : mv_polynomial σ R) -lemma coeff_weighted_homogeneous_component (d : σ →₀ ℕ) : +lemma coeff_weighted_homogeneous_component [decidable_eq M] (d : σ →₀ ℕ) : coeff d (weighted_homogeneous_component w n φ) = if weighted_degree' w d = n then coeff d φ else 0 := finsupp.filter_apply (λ d : σ →₀ ℕ, weighted_degree' w d = n) φ d -lemma weighted_homogeneous_component_apply : +lemma weighted_homogeneous_component_apply [decidable_eq M] : weighted_homogeneous_component w n φ = ∑ d in φ.support.filter (λ d, weighted_degree' w d = n), monomial d (coeff d φ) := finsupp.filter_eq_sum (λ d : σ →₀ ℕ, weighted_degree' w d = n) φ @@ -331,6 +333,7 @@ weighted degree `n`. -/ lemma weighted_homogeneous_component_is_weighted_homogeneous : (weighted_homogeneous_component w n φ).is_weighted_homogeneous w n := begin + classical, intros d hd, contrapose! hd, rw [coeff_weighted_homogeneous_component, if_neg hd] @@ -343,6 +346,7 @@ by simp only [C_mul', linear_map.map_smul] lemma weighted_homogeneous_component_eq_zero' (h : ∀ d : σ →₀ ℕ, d ∈ φ.support → weighted_degree' w d ≠ n) : weighted_homogeneous_component w n φ = 0 := begin + classical, rw [weighted_homogeneous_component_apply, sum_eq_zero], intros d hd, rw mem_filter at hd, exfalso, exact h _ hd.1 hd.2 @@ -351,6 +355,7 @@ end lemma weighted_homogeneous_component_eq_zero [semilattice_sup M] [order_bot M] (h : weighted_total_degree w φ < n) : weighted_homogeneous_component w n φ = 0 := begin + classical, rw [weighted_homogeneous_component_apply, sum_eq_zero], intros d hd, rw mem_filter at hd, exfalso, @@ -378,6 +383,7 @@ variable (w) lemma sum_weighted_homogeneous_component : finsum (λ m, weighted_homogeneous_component w m φ) = φ := begin + classical, rw finsum_eq_sum _ (weighted_homogeneous_component_finsupp φ), ext1 d, simp only [coeff_sum, coeff_weighted_homogeneous_component], @@ -386,7 +392,7 @@ begin { intros m hm hm', rw if_neg hm'.symm, }, { intro hm, rw if_pos rfl, simp only [finite.mem_to_finset, mem_support, ne.def, not_not] at hm, - have := coeff_weighted_homogeneous_component _ φ d, + have := coeff_weighted_homogeneous_component (_ : M) φ d, rw [hm, if_pos rfl, coeff_zero] at this, exact this.symm, }, end @@ -394,7 +400,7 @@ end variable {w} /-- The weighted homogeneous components of a weighted homogeneous polynomial. -/ -lemma weighted_homogeneous_component_weighted_homogeneous_polynomial (m n : M) +lemma weighted_homogeneous_component_weighted_homogeneous_polynomial [decidable_eq M] (m n : M) (p : mv_polynomial σ R) (h : p ∈ weighted_homogeneous_submodule R w n) : weighted_homogeneous_component w m p = if m = n then p else 0 := begin @@ -424,6 +430,7 @@ variables [canonically_ordered_add_monoid M] {w : σ → M} (φ : mv_polynomial @[simp] lemma weighted_homogeneous_component_zero [no_zero_smul_divisors ℕ M] (hw : ∀ i : σ, w i ≠ 0) : weighted_homogeneous_component w 0 φ = C (coeff 0 φ) := begin + classical, ext1 d, rcases em (d = 0) with (rfl|hd), { simp only [coeff_weighted_homogeneous_component, if_pos, map_zero, coeff_zero_C] }, diff --git a/src/ring_theory/witt_vector/mul_coeff.lean b/src/ring_theory/witt_vector/mul_coeff.lean index 196ef3ca87462..9903f792a5714 100644 --- a/src/ring_theory/witt_vector/mul_coeff.lean +++ b/src/ring_theory/witt_vector/mul_coeff.lean @@ -52,8 +52,8 @@ lemma witt_poly_prod_vars (n : ℕ) : (witt_poly_prod p n).vars ⊆ univ ×ˢ ra begin rw [witt_poly_prod], apply subset.trans (vars_mul _ _), - apply union_subset; - { apply subset.trans (vars_rename _ _), + refine union_subset _ _; + { refine subset.trans (vars_rename _ _) _, simp [witt_polynomial_vars,image_subset_iff] } end @@ -65,11 +65,11 @@ lemma witt_poly_prod_remainder_vars (n : ℕ) : (witt_poly_prod_remainder p n).vars ⊆ univ ×ˢ range n := begin rw [witt_poly_prod_remainder], - apply subset.trans (vars_sum_subset _ _), + refine subset.trans (vars_sum_subset _ _) _, rw bUnion_subset, intros x hx, apply subset.trans (vars_mul _ _), - apply union_subset, + refine union_subset _ _, { apply subset.trans (vars_pow _ _), have : (p : mv_polynomial (fin 2 × ℕ) ℤ) = (C (p : ℤ)), { simp only [int.cast_coe_nat, eq_int_cast] }, @@ -103,8 +103,8 @@ lemma remainder_vars (n : ℕ) : (remainder p n).vars ⊆ univ ×ˢ range (n + 1 begin rw [remainder], apply subset.trans (vars_mul _ _), - apply union_subset; - { apply subset.trans (vars_sum_subset _ _), + refine union_subset _ _; + { refine subset.trans (vars_sum_subset _ _) _, rw bUnion_subset, intros x hx, rw [rename_monomial, vars_monomial, finsupp.map_domain_single], @@ -207,7 +207,7 @@ lemma mul_poly_of_interest_vars (n : ℕ) : begin rw mul_poly_of_interest_aux5, apply subset.trans (vars_sub_subset _ _), - apply union_subset, + refine union_subset _ _, { apply remainder_vars }, { apply witt_poly_prod_remainder_vars } end From bf6a01357ff5684b1ebcd0f1a13be314fc82c0bf Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Mon, 15 May 2023 04:39:37 +0000 Subject: [PATCH 1008/1291] chore(measure_theory/{constructions/borel_space,integral/lebesgue}): split (#19015) Remove dependence on ```text import analysis.complex.basic import analysis.normed_space.basic import topology.instances.add_circle import topology.metric_space.metrizable ``` from `measure_theory.constructions.borel_space` (in particular this gets rid of the `operator_norm` dependency, which is currently the big blocker) and `measure_theory.integral.lebesgue`. --- src/analysis/calculus/fderiv_measurable.lean | 1 + src/dynamics/ergodic/conservative.lean | 2 +- .../basic.lean} | 257 +----------------- .../constructions/borel_space/complex.lean | 20 ++ .../borel_space/continuous_linear_map.lean | 94 +++++++ .../constructions/borel_space/metrizable.lean | 172 ++++++++++++ src/measure_theory/constructions/polish.lean | 2 +- src/measure_theory/covering/vitali.lean | 2 +- .../function/ae_measurable_order.lean | 2 +- src/measure_theory/function/ess_sup.lean | 2 +- src/measure_theory/function/floor.lean | 2 +- src/measure_theory/function/jacobian.lean | 1 + src/measure_theory/function/simple_func.lean | 2 +- .../function/special_functions/arctan.lean | 2 +- .../function/special_functions/basic.lean | 2 +- .../function/special_functions/inner.lean | 2 +- .../function/strongly_measurable/basic.lean | 3 +- src/measure_theory/integral/lebesgue.lean | 30 -- .../integral/lebesgue_normed_space.lean | 43 +++ src/measure_theory/integral/periodic.lean | 9 + src/measure_theory/measure/hausdorff.lean | 2 +- src/measure_theory/measure/regular.lean | 2 +- src/measure_theory/measure/stieltjes.lean | 2 +- 23 files changed, 357 insertions(+), 299 deletions(-) rename src/measure_theory/constructions/{borel_space.lean => borel_space/basic.lean} (87%) create mode 100644 src/measure_theory/constructions/borel_space/complex.lean create mode 100644 src/measure_theory/constructions/borel_space/continuous_linear_map.lean create mode 100644 src/measure_theory/constructions/borel_space/metrizable.lean create mode 100644 src/measure_theory/integral/lebesgue_normed_space.lean diff --git a/src/analysis/calculus/fderiv_measurable.lean b/src/analysis/calculus/fderiv_measurable.lean index 51d4260b59aff..ed414fa04a82b 100644 --- a/src/analysis/calculus/fderiv_measurable.lean +++ b/src/analysis/calculus/fderiv_measurable.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel, Yury Kudryashov -/ import analysis.calculus.deriv +import measure_theory.constructions.borel_space.continuous_linear_map import measure_theory.function.strongly_measurable.basic /-! diff --git a/src/dynamics/ergodic/conservative.lean b/src/dynamics/ergodic/conservative.lean index 94280422b3db5..c2a2fb13233c4 100644 --- a/src/dynamics/ergodic/conservative.lean +++ b/src/dynamics/ergodic/conservative.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ -import measure_theory.constructions.borel_space +import measure_theory.constructions.borel_space.basic import dynamics.ergodic.measure_preserving import combinatorics.pigeonhole diff --git a/src/measure_theory/constructions/borel_space.lean b/src/measure_theory/constructions/borel_space/basic.lean similarity index 87% rename from src/measure_theory/constructions/borel_space.lean rename to src/measure_theory/constructions/borel_space/basic.lean index e8067d5ceca54..c52b04a8be156 100644 --- a/src/measure_theory/constructions/borel_space.lean +++ b/src/measure_theory/constructions/borel_space/basic.lean @@ -3,20 +3,18 @@ Copyright (c) 2017 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Yury Kudryashov -/ -import analysis.complex.basic -import analysis.normed_space.finite_dimension +import analysis.normed.group.basic import measure_theory.function.ae_measurable_sequence import measure_theory.group.arithmetic import measure_theory.lattice import measure_theory.measure.open_pos import topology.algebra.order.liminf_limsup import topology.continuous_function.basic -import topology.instances.add_circle import topology.instances.ereal +import topology.metric_space.hausdorff_distance import topology.G_delta import topology.order.lattice import topology.semicontinuous -import topology.metric_space.metrizable /-! # Borel (measurable) space @@ -1327,11 +1325,6 @@ instance nat.borel_space : borel_space ℕ := ⟨borel_eq_top_of_discrete.symm instance int.borel_space : borel_space ℤ := ⟨borel_eq_top_of_discrete.symm⟩ instance rat.borel_space : borel_space ℚ := ⟨borel_eq_top_of_countable.symm⟩ -@[priority 900] -instance is_R_or_C.measurable_space {𝕜 : Type*} [is_R_or_C 𝕜] : measurable_space 𝕜 := borel 𝕜 -@[priority 900] -instance is_R_or_C.borel_space {𝕜 : Type*} [is_R_or_C 𝕜] : borel_space 𝕜 := ⟨rfl⟩ - /- Instances on `real` and `complex` are special cases of `is_R_or_C` but without these instances, Lean fails to prove `borel_space (ι → ℝ)`, so we leave them here. -/ @@ -1347,18 +1340,6 @@ instance ennreal.borel_space : borel_space ℝ≥0∞ := ⟨rfl⟩ instance ereal.measurable_space : measurable_space ereal := borel ereal instance ereal.borel_space : borel_space ereal := ⟨rfl⟩ -instance complex.measurable_space : measurable_space ℂ := borel ℂ -instance complex.borel_space : borel_space ℂ := ⟨rfl⟩ - -instance add_circle.measurable_space {a : ℝ} : measurable_space (add_circle a) := -borel (add_circle a) - -instance add_circle.borel_space {a : ℝ} : borel_space (add_circle a) := ⟨rfl⟩ - -@[measurability] protected lemma add_circle.measurable_mk' {a : ℝ} : - measurable (coe : ℝ → add_circle a) := -continuous.measurable $ add_circle.continuous_mk' a - /-- One can cut out `ℝ≥0∞` into the sets `{0}`, `Ico (t^n) (t^(n+1))` for `n : ℤ` and `{∞}`. This gives a way to compute the measure of a set in terms of sets on which a given function `f` does not fluctuate by more than `t`. -/ @@ -1899,237 +1880,3 @@ lemma ae_measurable.ennnorm {f : β → α} {μ : measure β} (hf : ae_measurabl measurable_ennnorm.comp_ae_measurable hf end normed_add_comm_group - -section limits - -variables [topological_space β] [pseudo_metrizable_space β] [measurable_space β] [borel_space β] - -open metric - -/-- A limit (over a general filter) of measurable `ℝ≥0∞` valued functions is measurable. -/ -lemma measurable_of_tendsto_ennreal' {ι} {f : ι → α → ℝ≥0∞} {g : α → ℝ≥0∞} (u : filter ι) - [ne_bot u] [is_countably_generated u] (hf : ∀ i, measurable (f i)) (lim : tendsto f u (𝓝 g)) : - measurable g := -begin - rcases u.exists_seq_tendsto with ⟨x, hx⟩, - rw [tendsto_pi_nhds] at lim, - have : (λ y, liminf (λ n, (f (x n) y : ℝ≥0∞)) at_top) = g := - by { ext1 y, exact ((lim y).comp hx).liminf_eq, }, - rw ← this, - show measurable (λ y, liminf (λ n, (f (x n) y : ℝ≥0∞)) at_top), - exact measurable_liminf (λ n, hf (x n)), -end - -/-- A sequential limit of measurable `ℝ≥0∞` valued functions is measurable. -/ -lemma measurable_of_tendsto_ennreal {f : ℕ → α → ℝ≥0∞} {g : α → ℝ≥0∞} - (hf : ∀ i, measurable (f i)) (lim : tendsto f at_top (𝓝 g)) : measurable g := -measurable_of_tendsto_ennreal' at_top hf lim - -/-- A limit (over a general filter) of measurable `ℝ≥0` valued functions is measurable. -/ -lemma measurable_of_tendsto_nnreal' {ι} {f : ι → α → ℝ≥0} {g : α → ℝ≥0} (u : filter ι) - [ne_bot u] [is_countably_generated u] (hf : ∀ i, measurable (f i)) (lim : tendsto f u (𝓝 g)) : - measurable g := -begin - simp_rw [← measurable_coe_nnreal_ennreal_iff] at hf ⊢, - refine measurable_of_tendsto_ennreal' u hf _, - rw tendsto_pi_nhds at lim ⊢, - exact λ x, (ennreal.continuous_coe.tendsto (g x)).comp (lim x), -end - -/-- A sequential limit of measurable `ℝ≥0` valued functions is measurable. -/ -lemma measurable_of_tendsto_nnreal {f : ℕ → α → ℝ≥0} {g : α → ℝ≥0} - (hf : ∀ i, measurable (f i)) (lim : tendsto f at_top (𝓝 g)) : measurable g := -measurable_of_tendsto_nnreal' at_top hf lim - -/-- A limit (over a general filter) of measurable functions valued in a (pseudo) metrizable space is -measurable. -/ -lemma measurable_of_tendsto_metrizable' {ι} {f : ι → α → β} {g : α → β} - (u : filter ι) [ne_bot u] [is_countably_generated u] - (hf : ∀ i, measurable (f i)) (lim : tendsto f u (𝓝 g)) : - measurable g := -begin - letI : pseudo_metric_space β := pseudo_metrizable_space_pseudo_metric β, - apply measurable_of_is_closed', intros s h1s h2s h3s, - have : measurable (λ x, inf_nndist (g x) s), - { suffices : tendsto (λ i x, inf_nndist (f i x) s) u (𝓝 (λ x, inf_nndist (g x) s)), - from measurable_of_tendsto_nnreal' u (λ i, (hf i).inf_nndist) this, - rw [tendsto_pi_nhds] at lim ⊢, intro x, - exact ((continuous_inf_nndist_pt s).tendsto (g x)).comp (lim x) }, - have h4s : g ⁻¹' s = (λ x, inf_nndist (g x) s) ⁻¹' {0}, - { ext x, simp [h1s, ← h1s.mem_iff_inf_dist_zero h2s, ← nnreal.coe_eq_zero] }, - rw [h4s], exact this (measurable_set_singleton 0), -end - -/-- A sequential limit of measurable functions valued in a (pseudo) metrizable space is -measurable. -/ -lemma measurable_of_tendsto_metrizable {f : ℕ → α → β} {g : α → β} - (hf : ∀ i, measurable (f i)) (lim : tendsto f at_top (𝓝 g)) : - measurable g := -measurable_of_tendsto_metrizable' at_top hf lim - -lemma ae_measurable_of_tendsto_metrizable_ae {ι} - {μ : measure α} {f : ι → α → β} {g : α → β} - (u : filter ι) [hu : ne_bot u] [is_countably_generated u] - (hf : ∀ n, ae_measurable (f n) μ) (h_tendsto : ∀ᵐ x ∂μ, tendsto (λ n, f n x) u (𝓝 (g x))) : - ae_measurable g μ := -begin - rcases u.exists_seq_tendsto with ⟨v, hv⟩, - have h'f : ∀ n, ae_measurable (f (v n)) μ := λ n, hf (v n), - set p : α → (ℕ → β) → Prop := λ x f', tendsto (λ n, f' n) at_top (𝓝 (g x)), - have hp : ∀ᵐ x ∂μ, p x (λ n, f (v n) x), - by filter_upwards [h_tendsto] with x hx using hx.comp hv, - set ae_seq_lim := λ x, ite (x ∈ ae_seq_set h'f p) (g x) (⟨f (v 0) x⟩ : nonempty β).some with hs, - refine ⟨ae_seq_lim, measurable_of_tendsto_metrizable' at_top (ae_seq.measurable h'f p) - (tendsto_pi_nhds.mpr (λ x, _)), _⟩, - { simp_rw [ae_seq, ae_seq_lim], - split_ifs with hx, - { simp_rw ae_seq.mk_eq_fun_of_mem_ae_seq_set h'f hx, - exact @ae_seq.fun_prop_of_mem_ae_seq_set _ α β _ _ _ _ _ h'f x hx, }, - { exact tendsto_const_nhds } }, - { exact (ite_ae_eq_of_measure_compl_zero g (λ x, (⟨f (v 0) x⟩ : nonempty β).some) - (ae_seq_set h'f p) (ae_seq.measure_compl_ae_seq_set_eq_zero h'f hp)).symm }, -end - -lemma ae_measurable_of_tendsto_metrizable_ae' {μ : measure α} {f : ℕ → α → β} {g : α → β} - (hf : ∀ n, ae_measurable (f n) μ) - (h_ae_tendsto : ∀ᵐ x ∂μ, tendsto (λ n, f n x) at_top (𝓝 (g x))) : - ae_measurable g μ := -ae_measurable_of_tendsto_metrizable_ae at_top hf h_ae_tendsto - -lemma ae_measurable_of_unif_approx {β} [measurable_space β] [pseudo_metric_space β] [borel_space β] - {μ : measure α} {g : α → β} - (hf : ∀ ε > (0 : ℝ), ∃ (f : α → β), ae_measurable f μ ∧ ∀ᵐ x ∂μ, dist (f x) (g x) ≤ ε) : - ae_measurable g μ := -begin - obtain ⟨u, u_anti, u_pos, u_lim⟩ : - ∃ (u : ℕ → ℝ), strict_anti u ∧ (∀ (n : ℕ), 0 < u n) ∧ tendsto u at_top (𝓝 0) := - exists_seq_strict_anti_tendsto (0 : ℝ), - choose f Hf using λ (n : ℕ), hf (u n) (u_pos n), - have : ∀ᵐ x ∂μ, tendsto (λ n, f n x) at_top (𝓝 (g x)), - { have : ∀ᵐ x ∂ μ, ∀ n, dist (f n x) (g x) ≤ u n := ae_all_iff.2 (λ n, (Hf n).2), - filter_upwards [this], - assume x hx, - rw tendsto_iff_dist_tendsto_zero, - exact squeeze_zero (λ n, dist_nonneg) hx u_lim }, - exact ae_measurable_of_tendsto_metrizable_ae' (λ n, (Hf n).1) this, -end - -lemma measurable_of_tendsto_metrizable_ae {μ : measure α} [μ.is_complete] {f : ℕ → α → β} - {g : α → β} (hf : ∀ n, measurable (f n)) - (h_ae_tendsto : ∀ᵐ x ∂μ, tendsto (λ n, f n x) at_top (𝓝 (g x))) : - measurable g := -ae_measurable_iff_measurable.mp - (ae_measurable_of_tendsto_metrizable_ae' (λ i, (hf i).ae_measurable) h_ae_tendsto) - -lemma measurable_limit_of_tendsto_metrizable_ae {ι} [countable ι] [nonempty ι] {μ : measure α} - {f : ι → α → β} {L : filter ι} [L.is_countably_generated] (hf : ∀ n, ae_measurable (f n) μ) - (h_ae_tendsto : ∀ᵐ x ∂μ, ∃ l : β, tendsto (λ n, f n x) L (𝓝 l)) : - ∃ (f_lim : α → β) (hf_lim_meas : measurable f_lim), - ∀ᵐ x ∂μ, tendsto (λ n, f n x) L (𝓝 (f_lim x)) := -begin - inhabit ι, - unfreezingI { rcases eq_or_ne L ⊥ with rfl | hL }, - { exact ⟨(hf default).mk _, (hf default).measurable_mk, - eventually_of_forall $ λ x, tendsto_bot⟩ }, - haveI : ne_bot L := ⟨hL⟩, - let p : α → (ι → β) → Prop := λ x f', ∃ l : β, tendsto (λ n, f' n) L (𝓝 l), - have hp_mem : ∀ x ∈ ae_seq_set hf p, p x (λ n, f n x), - from λ x hx, ae_seq.fun_prop_of_mem_ae_seq_set hf hx, - have h_ae_eq : ∀ᵐ x ∂μ, ∀ n, ae_seq hf p n x = f n x, - from ae_seq.ae_seq_eq_fun_ae hf h_ae_tendsto, - let f_lim : α → β := λ x, dite (x ∈ ae_seq_set hf p) (λ h, (hp_mem x h).some) - (λ h, (⟨f default x⟩ : nonempty β).some), - have hf_lim : ∀ x, tendsto (λ n, ae_seq hf p n x) L (𝓝 (f_lim x)), - { intros x, - simp only [f_lim, ae_seq], - split_ifs, - { refine (hp_mem x h).some_spec.congr (λ n, _), - exact (ae_seq.mk_eq_fun_of_mem_ae_seq_set hf h n).symm }, - { exact tendsto_const_nhds, }, }, - have h_ae_tendsto_f_lim : ∀ᵐ x ∂μ, tendsto (λ n, f n x) L (𝓝 (f_lim x)), - from h_ae_eq.mono (λ x hx, (hf_lim x).congr hx), - have h_f_lim_meas : measurable f_lim, - from measurable_of_tendsto_metrizable' L (ae_seq.measurable hf p) - (tendsto_pi_nhds.mpr (λ x, hf_lim x)), - exact ⟨f_lim, h_f_lim_meas, h_ae_tendsto_f_lim⟩, -end - -end limits - -namespace continuous_linear_map - -variables {𝕜 : Type*} [normed_field 𝕜] -variables {E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] [measurable_space E] - [opens_measurable_space E] {F : Type*} [normed_add_comm_group F] [normed_space 𝕜 F] - [measurable_space F] [borel_space F] - -@[measurability] -protected lemma measurable (L : E →L[𝕜] F) : measurable L := -L.continuous.measurable - -lemma measurable_comp (L : E →L[𝕜] F) {φ : α → E} (φ_meas : measurable φ) : - measurable (λ (a : α), L (φ a)) := -L.measurable.comp φ_meas - -end continuous_linear_map - -namespace continuous_linear_map - -variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] -variables {E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] - {F : Type*} [normed_add_comm_group F] [normed_space 𝕜 F] - -instance : measurable_space (E →L[𝕜] F) := borel _ - -instance : borel_space (E →L[𝕜] F) := ⟨rfl⟩ - -@[measurability] -lemma measurable_apply [measurable_space F] [borel_space F] (x : E) : - measurable (λ f : E →L[𝕜] F, f x) := -(apply 𝕜 F x).continuous.measurable - -@[measurability] -lemma measurable_apply' [measurable_space E] [opens_measurable_space E] - [measurable_space F] [borel_space F] : - measurable (λ (x : E) (f : E →L[𝕜] F), f x) := -measurable_pi_lambda _ $ λ f, f.measurable - -@[measurability] -lemma measurable_coe [measurable_space F] [borel_space F] : - measurable (λ (f : E →L[𝕜] F) (x : E), f x) := -measurable_pi_lambda _ measurable_apply - -end continuous_linear_map - -section continuous_linear_map_nontrivially_normed_field - -variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] -variables {E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] [measurable_space E] - [borel_space E] {F : Type*} [normed_add_comm_group F] [normed_space 𝕜 F] - -@[measurability] -lemma measurable.apply_continuous_linear_map {φ : α → F →L[𝕜] E} (hφ : measurable φ) (v : F) : - measurable (λ a, φ a v) := -(continuous_linear_map.apply 𝕜 E v).measurable.comp hφ - -@[measurability] -lemma ae_measurable.apply_continuous_linear_map {φ : α → F →L[𝕜] E} {μ : measure α} - (hφ : ae_measurable φ μ) (v : F) : ae_measurable (λ a, φ a v) μ := -(continuous_linear_map.apply 𝕜 E v).measurable.comp_ae_measurable hφ - -end continuous_linear_map_nontrivially_normed_field - -section normed_space -variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] [complete_space 𝕜] [measurable_space 𝕜] -variables [borel_space 𝕜] {E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] - [measurable_space E] [borel_space E] - -lemma measurable_smul_const {f : α → 𝕜} {c : E} (hc : c ≠ 0) : - measurable (λ x, f x • c) ↔ measurable f := -(closed_embedding_smul_left hc).measurable_embedding.measurable_comp_iff - -lemma ae_measurable_smul_const {f : α → 𝕜} {μ : measure α} {c : E} (hc : c ≠ 0) : - ae_measurable (λ x, f x • c) μ ↔ ae_measurable f μ := -(closed_embedding_smul_left hc).measurable_embedding.ae_measurable_comp_iff - -end normed_space diff --git a/src/measure_theory/constructions/borel_space/complex.lean b/src/measure_theory/constructions/borel_space/complex.lean new file mode 100644 index 0000000000000..1a69089f7ab98 --- /dev/null +++ b/src/measure_theory/constructions/borel_space/complex.lean @@ -0,0 +1,20 @@ +/- +Copyright (c) 2020 Yury Kudryashov. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yury Kudryashov +-/ +import analysis.complex.basic +import measure_theory.constructions.borel_space.basic + +/-! # Equip `ℂ` with the Borel sigma-algebra -/ + +noncomputable theory + +@[priority 900] +instance is_R_or_C.measurable_space {𝕜 : Type*} [is_R_or_C 𝕜] : measurable_space 𝕜 := borel 𝕜 +@[priority 900] +instance is_R_or_C.borel_space {𝕜 : Type*} [is_R_or_C 𝕜] : borel_space 𝕜 := ⟨rfl⟩ + + +instance complex.measurable_space : measurable_space ℂ := borel ℂ +instance complex.borel_space : borel_space ℂ := ⟨rfl⟩ diff --git a/src/measure_theory/constructions/borel_space/continuous_linear_map.lean b/src/measure_theory/constructions/borel_space/continuous_linear_map.lean new file mode 100644 index 0000000000000..6d7908b94a05a --- /dev/null +++ b/src/measure_theory/constructions/borel_space/continuous_linear_map.lean @@ -0,0 +1,94 @@ +/- +Copyright (c) 2020 Patrick Massot. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Patrick Massot +-/ +import analysis.normed_space.finite_dimension +import measure_theory.constructions.borel_space.basic + +/-! +# Measurable functions in normed spaces + +-/ + +open measure_theory + +variables {α : Type*} [measurable_space α] + +namespace continuous_linear_map + +variables {𝕜 : Type*} [normed_field 𝕜] +variables {E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] [measurable_space E] + [opens_measurable_space E] {F : Type*} [normed_add_comm_group F] [normed_space 𝕜 F] + [measurable_space F] [borel_space F] + +@[measurability] +protected lemma measurable (L : E →L[𝕜] F) : measurable L := +L.continuous.measurable + +lemma measurable_comp (L : E →L[𝕜] F) {φ : α → E} (φ_meas : measurable φ) : + measurable (λ (a : α), L (φ a)) := +L.measurable.comp φ_meas + +end continuous_linear_map + +namespace continuous_linear_map + +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] +variables {E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] + {F : Type*} [normed_add_comm_group F] [normed_space 𝕜 F] + +instance : measurable_space (E →L[𝕜] F) := borel _ + +instance : borel_space (E →L[𝕜] F) := ⟨rfl⟩ + +@[measurability] +lemma measurable_apply [measurable_space F] [borel_space F] (x : E) : + measurable (λ f : E →L[𝕜] F, f x) := +(apply 𝕜 F x).continuous.measurable + +@[measurability] +lemma measurable_apply' [measurable_space E] [opens_measurable_space E] + [measurable_space F] [borel_space F] : + measurable (λ (x : E) (f : E →L[𝕜] F), f x) := +measurable_pi_lambda _ $ λ f, f.measurable + +@[measurability] +lemma measurable_coe [measurable_space F] [borel_space F] : + measurable (λ (f : E →L[𝕜] F) (x : E), f x) := +measurable_pi_lambda _ measurable_apply + +end continuous_linear_map + +section continuous_linear_map_nontrivially_normed_field + +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] +variables {E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] [measurable_space E] + [borel_space E] {F : Type*} [normed_add_comm_group F] [normed_space 𝕜 F] + +@[measurability] +lemma measurable.apply_continuous_linear_map {φ : α → F →L[𝕜] E} (hφ : measurable φ) (v : F) : + measurable (λ a, φ a v) := +(continuous_linear_map.apply 𝕜 E v).measurable.comp hφ + +@[measurability] +lemma ae_measurable.apply_continuous_linear_map {φ : α → F →L[𝕜] E} {μ : measure α} + (hφ : ae_measurable φ μ) (v : F) : ae_measurable (λ a, φ a v) μ := +(continuous_linear_map.apply 𝕜 E v).measurable.comp_ae_measurable hφ + +end continuous_linear_map_nontrivially_normed_field + +section normed_space +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] [complete_space 𝕜] [measurable_space 𝕜] +variables [borel_space 𝕜] {E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] + [measurable_space E] [borel_space E] + +lemma measurable_smul_const {f : α → 𝕜} {c : E} (hc : c ≠ 0) : + measurable (λ x, f x • c) ↔ measurable f := +(closed_embedding_smul_left hc).measurable_embedding.measurable_comp_iff + +lemma ae_measurable_smul_const {f : α → 𝕜} {μ : measure α} {c : E} (hc : c ≠ 0) : + ae_measurable (λ x, f x • c) μ ↔ ae_measurable f μ := +(closed_embedding_smul_left hc).measurable_embedding.ae_measurable_comp_iff + +end normed_space diff --git a/src/measure_theory/constructions/borel_space/metrizable.lean b/src/measure_theory/constructions/borel_space/metrizable.lean new file mode 100644 index 0000000000000..6572f6f437674 --- /dev/null +++ b/src/measure_theory/constructions/borel_space/metrizable.lean @@ -0,0 +1,172 @@ +/- +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 +-/ +import measure_theory.constructions.borel_space.basic +import topology.metric_space.metrizable + +/-! +# Measurable functions in (pseudo-)metrizable Borel spaces +-/ + +open filter measure_theory topological_space +open_locale classical topology nnreal ennreal measure_theory + +variables {α β : Type*} [measurable_space α] + +section limits + +variables [topological_space β] [pseudo_metrizable_space β] [measurable_space β] [borel_space β] + +open metric + +/-- A limit (over a general filter) of measurable `ℝ≥0∞` valued functions is measurable. -/ +lemma measurable_of_tendsto_ennreal' {ι} {f : ι → α → ℝ≥0∞} {g : α → ℝ≥0∞} (u : filter ι) + [ne_bot u] [is_countably_generated u] (hf : ∀ i, measurable (f i)) (lim : tendsto f u (𝓝 g)) : + measurable g := +begin + rcases u.exists_seq_tendsto with ⟨x, hx⟩, + rw [tendsto_pi_nhds] at lim, + have : (λ y, liminf (λ n, (f (x n) y : ℝ≥0∞)) at_top) = g := + by { ext1 y, exact ((lim y).comp hx).liminf_eq, }, + rw ← this, + show measurable (λ y, liminf (λ n, (f (x n) y : ℝ≥0∞)) at_top), + exact measurable_liminf (λ n, hf (x n)), +end + +/-- A sequential limit of measurable `ℝ≥0∞` valued functions is measurable. -/ +lemma measurable_of_tendsto_ennreal {f : ℕ → α → ℝ≥0∞} {g : α → ℝ≥0∞} + (hf : ∀ i, measurable (f i)) (lim : tendsto f at_top (𝓝 g)) : measurable g := +measurable_of_tendsto_ennreal' at_top hf lim + +/-- A limit (over a general filter) of measurable `ℝ≥0` valued functions is measurable. -/ +lemma measurable_of_tendsto_nnreal' {ι} {f : ι → α → ℝ≥0} {g : α → ℝ≥0} (u : filter ι) + [ne_bot u] [is_countably_generated u] (hf : ∀ i, measurable (f i)) (lim : tendsto f u (𝓝 g)) : + measurable g := +begin + simp_rw [← measurable_coe_nnreal_ennreal_iff] at hf ⊢, + refine measurable_of_tendsto_ennreal' u hf _, + rw tendsto_pi_nhds at lim ⊢, + exact λ x, (ennreal.continuous_coe.tendsto (g x)).comp (lim x), +end + +/-- A sequential limit of measurable `ℝ≥0` valued functions is measurable. -/ +lemma measurable_of_tendsto_nnreal {f : ℕ → α → ℝ≥0} {g : α → ℝ≥0} + (hf : ∀ i, measurable (f i)) (lim : tendsto f at_top (𝓝 g)) : measurable g := +measurable_of_tendsto_nnreal' at_top hf lim + +/-- A limit (over a general filter) of measurable functions valued in a (pseudo) metrizable space is +measurable. -/ +lemma measurable_of_tendsto_metrizable' {ι} {f : ι → α → β} {g : α → β} + (u : filter ι) [ne_bot u] [is_countably_generated u] + (hf : ∀ i, measurable (f i)) (lim : tendsto f u (𝓝 g)) : + measurable g := +begin + letI : pseudo_metric_space β := pseudo_metrizable_space_pseudo_metric β, + apply measurable_of_is_closed', intros s h1s h2s h3s, + have : measurable (λ x, inf_nndist (g x) s), + { suffices : tendsto (λ i x, inf_nndist (f i x) s) u (𝓝 (λ x, inf_nndist (g x) s)), + from measurable_of_tendsto_nnreal' u (λ i, (hf i).inf_nndist) this, + rw [tendsto_pi_nhds] at lim ⊢, intro x, + exact ((continuous_inf_nndist_pt s).tendsto (g x)).comp (lim x) }, + have h4s : g ⁻¹' s = (λ x, inf_nndist (g x) s) ⁻¹' {0}, + { ext x, simp [h1s, ← h1s.mem_iff_inf_dist_zero h2s, ← nnreal.coe_eq_zero] }, + rw [h4s], exact this (measurable_set_singleton 0), +end + +/-- A sequential limit of measurable functions valued in a (pseudo) metrizable space is +measurable. -/ +lemma measurable_of_tendsto_metrizable {f : ℕ → α → β} {g : α → β} + (hf : ∀ i, measurable (f i)) (lim : tendsto f at_top (𝓝 g)) : + measurable g := +measurable_of_tendsto_metrizable' at_top hf lim + +lemma ae_measurable_of_tendsto_metrizable_ae {ι} + {μ : measure α} {f : ι → α → β} {g : α → β} + (u : filter ι) [hu : ne_bot u] [is_countably_generated u] + (hf : ∀ n, ae_measurable (f n) μ) (h_tendsto : ∀ᵐ x ∂μ, tendsto (λ n, f n x) u (𝓝 (g x))) : + ae_measurable g μ := +begin + rcases u.exists_seq_tendsto with ⟨v, hv⟩, + have h'f : ∀ n, ae_measurable (f (v n)) μ := λ n, hf (v n), + set p : α → (ℕ → β) → Prop := λ x f', tendsto (λ n, f' n) at_top (𝓝 (g x)), + have hp : ∀ᵐ x ∂μ, p x (λ n, f (v n) x), + by filter_upwards [h_tendsto] with x hx using hx.comp hv, + set ae_seq_lim := λ x, ite (x ∈ ae_seq_set h'f p) (g x) (⟨f (v 0) x⟩ : nonempty β).some with hs, + refine ⟨ae_seq_lim, measurable_of_tendsto_metrizable' at_top (ae_seq.measurable h'f p) + (tendsto_pi_nhds.mpr (λ x, _)), _⟩, + { simp_rw [ae_seq, ae_seq_lim], + split_ifs with hx, + { simp_rw ae_seq.mk_eq_fun_of_mem_ae_seq_set h'f hx, + exact @ae_seq.fun_prop_of_mem_ae_seq_set _ α β _ _ _ _ _ h'f x hx, }, + { exact tendsto_const_nhds } }, + { exact (ite_ae_eq_of_measure_compl_zero g (λ x, (⟨f (v 0) x⟩ : nonempty β).some) + (ae_seq_set h'f p) (ae_seq.measure_compl_ae_seq_set_eq_zero h'f hp)).symm }, +end + +lemma ae_measurable_of_tendsto_metrizable_ae' {μ : measure α} {f : ℕ → α → β} {g : α → β} + (hf : ∀ n, ae_measurable (f n) μ) + (h_ae_tendsto : ∀ᵐ x ∂μ, tendsto (λ n, f n x) at_top (𝓝 (g x))) : + ae_measurable g μ := +ae_measurable_of_tendsto_metrizable_ae at_top hf h_ae_tendsto + +lemma ae_measurable_of_unif_approx {β} [measurable_space β] [pseudo_metric_space β] [borel_space β] + {μ : measure α} {g : α → β} + (hf : ∀ ε > (0 : ℝ), ∃ (f : α → β), ae_measurable f μ ∧ ∀ᵐ x ∂μ, dist (f x) (g x) ≤ ε) : + ae_measurable g μ := +begin + obtain ⟨u, u_anti, u_pos, u_lim⟩ : + ∃ (u : ℕ → ℝ), strict_anti u ∧ (∀ (n : ℕ), 0 < u n) ∧ tendsto u at_top (𝓝 0) := + exists_seq_strict_anti_tendsto (0 : ℝ), + choose f Hf using λ (n : ℕ), hf (u n) (u_pos n), + have : ∀ᵐ x ∂μ, tendsto (λ n, f n x) at_top (𝓝 (g x)), + { have : ∀ᵐ x ∂ μ, ∀ n, dist (f n x) (g x) ≤ u n := ae_all_iff.2 (λ n, (Hf n).2), + filter_upwards [this], + assume x hx, + rw tendsto_iff_dist_tendsto_zero, + exact squeeze_zero (λ n, dist_nonneg) hx u_lim }, + exact ae_measurable_of_tendsto_metrizable_ae' (λ n, (Hf n).1) this, +end + +lemma measurable_of_tendsto_metrizable_ae {μ : measure α} [μ.is_complete] {f : ℕ → α → β} + {g : α → β} (hf : ∀ n, measurable (f n)) + (h_ae_tendsto : ∀ᵐ x ∂μ, tendsto (λ n, f n x) at_top (𝓝 (g x))) : + measurable g := +ae_measurable_iff_measurable.mp + (ae_measurable_of_tendsto_metrizable_ae' (λ i, (hf i).ae_measurable) h_ae_tendsto) + +lemma measurable_limit_of_tendsto_metrizable_ae {ι} [countable ι] [nonempty ι] {μ : measure α} + {f : ι → α → β} {L : filter ι} [L.is_countably_generated] (hf : ∀ n, ae_measurable (f n) μ) + (h_ae_tendsto : ∀ᵐ x ∂μ, ∃ l : β, tendsto (λ n, f n x) L (𝓝 l)) : + ∃ (f_lim : α → β) (hf_lim_meas : measurable f_lim), + ∀ᵐ x ∂μ, tendsto (λ n, f n x) L (𝓝 (f_lim x)) := +begin + inhabit ι, + unfreezingI { rcases eq_or_ne L ⊥ with rfl | hL }, + { exact ⟨(hf default).mk _, (hf default).measurable_mk, + eventually_of_forall $ λ x, tendsto_bot⟩ }, + haveI : ne_bot L := ⟨hL⟩, + let p : α → (ι → β) → Prop := λ x f', ∃ l : β, tendsto (λ n, f' n) L (𝓝 l), + have hp_mem : ∀ x ∈ ae_seq_set hf p, p x (λ n, f n x), + from λ x hx, ae_seq.fun_prop_of_mem_ae_seq_set hf hx, + have h_ae_eq : ∀ᵐ x ∂μ, ∀ n, ae_seq hf p n x = f n x, + from ae_seq.ae_seq_eq_fun_ae hf h_ae_tendsto, + let f_lim : α → β := λ x, dite (x ∈ ae_seq_set hf p) (λ h, (hp_mem x h).some) + (λ h, (⟨f default x⟩ : nonempty β).some), + have hf_lim : ∀ x, tendsto (λ n, ae_seq hf p n x) L (𝓝 (f_lim x)), + { intros x, + simp only [f_lim, ae_seq], + split_ifs, + { refine (hp_mem x h).some_spec.congr (λ n, _), + exact (ae_seq.mk_eq_fun_of_mem_ae_seq_set hf h n).symm }, + { exact tendsto_const_nhds, }, }, + have h_ae_tendsto_f_lim : ∀ᵐ x ∂μ, tendsto (λ n, f n x) L (𝓝 (f_lim x)), + from h_ae_eq.mono (λ x hx, (hf_lim x).congr hx), + have h_f_lim_meas : measurable f_lim, + from measurable_of_tendsto_metrizable' L (ae_seq.measurable hf p) + (tendsto_pi_nhds.mpr (λ x, hf_lim x)), + exact ⟨f_lim, h_f_lim_meas, h_ae_tendsto_f_lim⟩, +end + +end limits diff --git a/src/measure_theory/constructions/polish.lean b/src/measure_theory/constructions/polish.lean index 546b85d2501a0..9c7d3638995c9 100644 --- a/src/measure_theory/constructions/polish.lean +++ b/src/measure_theory/constructions/polish.lean @@ -5,7 +5,7 @@ Authors: Sébastien Gouëzel, Felix Weilacher -/ import data.real.cardinality import topology.perfect -import measure_theory.constructions.borel_space +import measure_theory.constructions.borel_space.basic /-! # The Borel sigma-algebra on Polish spaces diff --git a/src/measure_theory/covering/vitali.lean b/src/measure_theory/covering/vitali.lean index 16df444849fe1..622c70dee1868 100644 --- a/src/measure_theory/covering/vitali.lean +++ b/src/measure_theory/covering/vitali.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ import topology.metric_space.basic -import measure_theory.constructions.borel_space +import measure_theory.constructions.borel_space.basic import measure_theory.covering.vitali_family /-! diff --git a/src/measure_theory/function/ae_measurable_order.lean b/src/measure_theory/function/ae_measurable_order.lean index 2c5766576668d..478ef9275ea72 100644 --- a/src/measure_theory/function/ae_measurable_order.lean +++ b/src/measure_theory/function/ae_measurable_order.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ -import measure_theory.constructions.borel_space +import measure_theory.constructions.borel_space.basic /-! # Measurability criterion for ennreal-valued functions diff --git a/src/measure_theory/function/ess_sup.lean b/src/measure_theory/function/ess_sup.lean index 23ecbb3325864..01c891dcc97ea 100644 --- a/src/measure_theory/function/ess_sup.lean +++ b/src/measure_theory/function/ess_sup.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Rémy Degenne. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ -import measure_theory.constructions.borel_space +import measure_theory.constructions.borel_space.basic import order.filter.ennreal /-! diff --git a/src/measure_theory/function/floor.lean b/src/measure_theory/function/floor.lean index 9d1199f835b65..41b72efe283a2 100644 --- a/src/measure_theory/function/floor.lean +++ b/src/measure_theory/function/floor.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Yury G. Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury G. Kudryashov -/ -import measure_theory.constructions.borel_space +import measure_theory.constructions.borel_space.basic /-! # Measurability of `⌊x⌋` etc diff --git a/src/measure_theory/function/jacobian.lean b/src/measure_theory/function/jacobian.lean index 3be38f43111b4..2055a245a40ea 100644 --- a/src/measure_theory/function/jacobian.lean +++ b/src/measure_theory/function/jacobian.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ +import measure_theory.constructions.borel_space.continuous_linear_map import measure_theory.covering.besicovitch_vector_space import measure_theory.measure.haar_lebesgue import analysis.normed_space.pointwise diff --git a/src/measure_theory/function/simple_func.lean b/src/measure_theory/function/simple_func.lean index 9051f119f2976..0c8816199c41d 100644 --- a/src/measure_theory/function/simple_func.lean +++ b/src/measure_theory/function/simple_func.lean @@ -3,7 +3,7 @@ Copyright (c) 2018 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Johannes Hölzl -/ -import measure_theory.constructions.borel_space +import measure_theory.constructions.borel_space.basic import algebra.indicator_function import algebra.support diff --git a/src/measure_theory/function/special_functions/arctan.lean b/src/measure_theory/function/special_functions/arctan.lean index 8817055039e83..d1fdd51f129b3 100644 --- a/src/measure_theory/function/special_functions/arctan.lean +++ b/src/measure_theory/function/special_functions/arctan.lean @@ -5,7 +5,7 @@ Authors: Yury Kudryashov -/ import analysis.special_functions.trigonometric.arctan -import measure_theory.constructions.borel_space +import measure_theory.constructions.borel_space.basic /-! # Measurability of arctan diff --git a/src/measure_theory/function/special_functions/basic.lean b/src/measure_theory/function/special_functions/basic.lean index e394cccf6cd95..78d900f0c4bb9 100644 --- a/src/measure_theory/function/special_functions/basic.lean +++ b/src/measure_theory/function/special_functions/basic.lean @@ -6,7 +6,7 @@ Authors: Yury Kudryashov import analysis.special_functions.pow import data.is_R_or_C.lemmas -import measure_theory.constructions.borel_space +import measure_theory.constructions.borel_space.complex /-! # Measurability of real and complex functions diff --git a/src/measure_theory/function/special_functions/inner.lean b/src/measure_theory/function/special_functions/inner.lean index 627872248de06..1f11270445b53 100644 --- a/src/measure_theory/function/special_functions/inner.lean +++ b/src/measure_theory/function/special_functions/inner.lean @@ -5,7 +5,7 @@ Authors: Yury Kudryashov -/ import analysis.inner_product_space.basic -import measure_theory.constructions.borel_space +import measure_theory.constructions.borel_space.complex /-! # Measurability of scalar products diff --git a/src/measure_theory/function/strongly_measurable/basic.lean b/src/measure_theory/function/strongly_measurable/basic.lean index 610bde59e6b99..df0f43cb089d0 100644 --- a/src/measure_theory/function/strongly_measurable/basic.lean +++ b/src/measure_theory/function/strongly_measurable/basic.lean @@ -3,8 +3,9 @@ Copyright (c) 2021 Rémy Degenne. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne, Sébastien Gouëzel -/ +import analysis.normed_space.finite_dimension import analysis.normed_space.bounded_linear_maps -import topology.metric_space.metrizable +import measure_theory.constructions.borel_space.metrizable import measure_theory.integral.lebesgue import measure_theory.function.simple_func_dense diff --git a/src/measure_theory/integral/lebesgue.lean b/src/measure_theory/integral/lebesgue.lean index 7f5ee430c2000..413dc0a066358 100644 --- a/src/measure_theory/integral/lebesgue.lean +++ b/src/measure_theory/integral/lebesgue.lean @@ -1684,36 +1684,6 @@ begin { exact hf (measurable_set_singleton 0).compl }, end -lemma ae_measurable_with_density_iff {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] - [topological_space.second_countable_topology E] [measurable_space E] [borel_space E] - {f : α → ℝ≥0} (hf : measurable f) {g : α → E} : - ae_measurable g (μ.with_density (λ x, (f x : ℝ≥0∞))) ↔ ae_measurable (λ x, (f x : ℝ) • g x) μ := -begin - split, - { rintros ⟨g', g'meas, hg'⟩, - have A : measurable_set {x : α | f x ≠ 0} := (hf (measurable_set_singleton 0)).compl, - refine ⟨λ x, (f x : ℝ) • g' x, hf.coe_nnreal_real.smul g'meas, _⟩, - apply @ae_of_ae_restrict_of_ae_restrict_compl _ _ _ {x | f x ≠ 0}, - { rw [eventually_eq, ae_with_density_iff hf.coe_nnreal_ennreal] at hg', - rw ae_restrict_iff' A, - filter_upwards [hg'], - assume a ha h'a, - have : (f a : ℝ≥0∞) ≠ 0, by simpa only [ne.def, coe_eq_zero] using h'a, - rw ha this }, - { filter_upwards [ae_restrict_mem A.compl], - assume x hx, - simp only [not_not, mem_set_of_eq, mem_compl_iff] at hx, - simp [hx] } }, - { rintros ⟨g', g'meas, hg'⟩, - refine ⟨λ x, (f x : ℝ)⁻¹ • g' x, hf.coe_nnreal_real.inv.smul g'meas, _⟩, - rw [eventually_eq, ae_with_density_iff hf.coe_nnreal_ennreal], - filter_upwards [hg'], - assume x hx h'x, - rw [← hx, smul_smul, _root_.inv_mul_cancel, one_smul], - simp only [ne.def, coe_eq_zero] at h'x, - simpa only [nnreal.coe_eq_zero, ne.def] using h'x } -end - lemma ae_measurable_with_density_ennreal_iff {f : α → ℝ≥0} (hf : measurable f) {g : α → ℝ≥0∞} : ae_measurable g (μ.with_density (λ x, (f x : ℝ≥0∞))) ↔ ae_measurable (λ x, (f x : ℝ≥0∞) * g x) μ := diff --git a/src/measure_theory/integral/lebesgue_normed_space.lean b/src/measure_theory/integral/lebesgue_normed_space.lean new file mode 100644 index 0000000000000..4301c8df6ba13 --- /dev/null +++ b/src/measure_theory/integral/lebesgue_normed_space.lean @@ -0,0 +1,43 @@ +/- +Copyright (c) 2022 Sébastien Gouëzel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sébastien Gouëzel +-/ +import measure_theory.integral.lebesgue +import analysis.normed_space.basic + +/-! # A lemma about measurability with density under scalar multiplication in normed spaces -/ + +open measure_theory filter ennreal set +open_locale nnreal ennreal +variables {α β γ δ : Type*} {m : measurable_space α} {μ : measure_theory.measure α} + +lemma ae_measurable_with_density_iff {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] + [topological_space.second_countable_topology E] [measurable_space E] [borel_space E] + {f : α → ℝ≥0} (hf : measurable f) {g : α → E} : + ae_measurable g (μ.with_density (λ x, (f x : ℝ≥0∞))) ↔ ae_measurable (λ x, (f x : ℝ) • g x) μ := +begin + split, + { rintros ⟨g', g'meas, hg'⟩, + have A : measurable_set {x : α | f x ≠ 0} := (hf (measurable_set_singleton 0)).compl, + refine ⟨λ x, (f x : ℝ) • g' x, hf.coe_nnreal_real.smul g'meas, _⟩, + apply @ae_of_ae_restrict_of_ae_restrict_compl _ _ _ {x | f x ≠ 0}, + { rw [eventually_eq, ae_with_density_iff hf.coe_nnreal_ennreal] at hg', + rw ae_restrict_iff' A, + filter_upwards [hg'], + assume a ha h'a, + have : (f a : ℝ≥0∞) ≠ 0, by simpa only [ne.def, coe_eq_zero] using h'a, + rw ha this }, + { filter_upwards [ae_restrict_mem A.compl], + assume x hx, + simp only [not_not, mem_set_of_eq, mem_compl_iff] at hx, + simp [hx] } }, + { rintros ⟨g', g'meas, hg'⟩, + refine ⟨λ x, (f x : ℝ)⁻¹ • g' x, hf.coe_nnreal_real.inv.smul g'meas, _⟩, + rw [eventually_eq, ae_with_density_iff hf.coe_nnreal_ennreal], + filter_upwards [hg'], + assume x hx h'x, + rw [← hx, smul_smul, _root_.inv_mul_cancel, one_smul], + simp only [ne.def, coe_eq_zero] at h'x, + simpa only [nnreal.coe_eq_zero, ne.def] using h'x } +end diff --git a/src/measure_theory/integral/periodic.lean b/src/measure_theory/integral/periodic.lean index 288be8f6d012a..48ed554e4831a 100644 --- a/src/measure_theory/integral/periodic.lean +++ b/src/measure_theory/integral/periodic.lean @@ -30,6 +30,15 @@ open_locale measure_theory nnreal ennreal local attribute [-instance] quotient_add_group.measurable_space quotient.measurable_space +noncomputable instance add_circle.measurable_space {a : ℝ} : measurable_space (add_circle a) := +borel (add_circle a) + +instance add_circle.borel_space {a : ℝ} : borel_space (add_circle a) := ⟨rfl⟩ + +@[measurability] protected lemma add_circle.measurable_mk' {a : ℝ} : + measurable (coe : ℝ → add_circle a) := +continuous.measurable $ add_circle.continuous_mk' a + lemma is_add_fundamental_domain_Ioc {T : ℝ} (hT : 0 < T) (t : ℝ) (μ : measure ℝ . volume_tac) : is_add_fundamental_domain (add_subgroup.zmultiples T) (Ioc t (t + T)) μ := begin diff --git a/src/measure_theory/measure/hausdorff.lean b/src/measure_theory/measure/hausdorff.lean index 2b28738a336d3..4c954fa1b605a 100644 --- a/src/measure_theory/measure/hausdorff.lean +++ b/src/measure_theory/measure/hausdorff.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import analysis.special_functions.pow -import measure_theory.constructions.borel_space +import measure_theory.constructions.borel_space.basic import measure_theory.measure.lebesgue import topology.metric_space.holder import topology.metric_space.metric_separated diff --git a/src/measure_theory/measure/regular.lean b/src/measure_theory/measure/regular.lean index a025ef6c8720b..4f3eb55b5584f 100644 --- a/src/measure_theory/measure/regular.lean +++ b/src/measure_theory/measure/regular.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris Van Doorn, Yury Kudryashov -/ -import measure_theory.constructions.borel_space +import measure_theory.constructions.borel_space.basic /-! # Regular measures diff --git a/src/measure_theory/measure/stieltjes.lean b/src/measure_theory/measure/stieltjes.lean index 9911b118cab92..07f34e0b845d4 100644 --- a/src/measure_theory/measure/stieltjes.lean +++ b/src/measure_theory/measure/stieltjes.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Yury Kudryashov, Sébastien Gouëzel -/ -import measure_theory.constructions.borel_space +import measure_theory.constructions.borel_space.basic import topology.algebra.order.left_right_lim /-! From 74a27133cf29446a0983779e37c8f829a85368f3 Mon Sep 17 00:00:00 2001 From: Michael Stoll Date: Mon, 15 May 2023 06:03:03 +0000 Subject: [PATCH 1009/1291] feat(number_theory/legendre_symbol/*): add some API for the Legendre and the Jacobi symbol (#19011) This PR adds some additional API lemmas for the Legendre symbol (related to solutions mod `p` of `x^2-a*y^2 = 0`) and for the Jacobi symbol. These are useful for a project formalizing the proof that a certain diophantine equation has no integral solutions, but also appear to make sense more generally. --- .../legendre_symbol/jacobi_symbol.lean | 44 +++++++++++++ .../quadratic_reciprocity.lean | 65 +++++++++++++++++++ 2 files changed, 109 insertions(+) diff --git a/src/number_theory/legendre_symbol/jacobi_symbol.lean b/src/number_theory/legendre_symbol/jacobi_symbol.lean index 9bbec60ab8148..597367d2409f4 100644 --- a/src/number_theory/legendre_symbol/jacobi_symbol.lean +++ b/src/number_theory/legendre_symbol/jacobi_symbol.lean @@ -201,6 +201,50 @@ end lemma mod_left' {a₁ a₂ : ℤ} {b : ℕ} (h : a₁ % b = a₂ % b) : J(a₁ | b) = J(a₂ | b) := by rw [mod_left, h, ← mod_left] +/-- If `p` is prime, `J(a | p) = -1` and `p` divides `x^2 - a*y^2`, then `p` must divide +`x` and `y`. -/ +lemma prime_dvd_of_eq_neg_one {p : ℕ} [fact p.prime] {a : ℤ} (h : J(a | p) = -1) + {x y : ℤ} (hxy : ↑p ∣ x ^ 2 - a * y ^ 2) : ↑p ∣ x ∧ ↑p ∣ y := +begin + rw [← legendre_sym.to_jacobi_sym] at h, + exact legendre_sym.prime_dvd_of_eq_neg_one h hxy, +end + +/-- We can pull out a product over a list in the first argument of the Jacobi symbol. -/ +lemma list_prod_left {l : list ℤ} {n : ℕ} : + J(l.prod | n) = (l.map (λ a, J(a | n))).prod := +begin + induction l with n l' ih, + { simp only [list.prod_nil, list.map_nil, one_left], }, + { rw [list.map, list.prod_cons, list.prod_cons, mul_left, ih], } +end + +/-- We can pull out a product over a list in the second argument of the Jacobi symbol. -/ +lemma list_prod_right {a : ℤ} {l : list ℕ} (hl : ∀ n ∈ l, n ≠ 0) : + J(a | l.prod) = (l.map (λ n, J(a | n))).prod := +begin + induction l with n l' ih, + { simp only [list.prod_nil, one_right, list.map_nil], }, + { have hn := hl n (list.mem_cons_self n l'), -- `n ≠ 0` + have hl' := list.prod_ne_zero (λ hf, hl 0 (list.mem_cons_of_mem _ hf) rfl), -- `l'.prod ≠ 0` + have h := λ m hm, hl m (list.mem_cons_of_mem _ hm), -- `∀ (m : ℕ), m ∈ l' → m ≠ 0` + rw [list.map, list.prod_cons, list.prod_cons, mul_right' a hn hl', ih h], } +end + +/-- If `J(a | n) = -1`, then `n` has a prime divisor `p` such that `J(a | p) = -1`. -/ +lemma eq_neg_one_at_prime_divisor_of_eq_neg_one {a : ℤ} {n : ℕ} (h : J(a | n) = -1) : + ∃ (p : ℕ) (hp : p.prime), p ∣ n ∧ J(a | p) = -1 := +begin + have hn₀ : n ≠ 0, + { rintro rfl, + rw [zero_right, eq_neg_self_iff] at h, + exact one_ne_zero h, }, + have hf₀ : ∀ p ∈ n.factors, p ≠ 0 := λ p hp, (nat.pos_of_mem_factors hp).ne.symm, + rw [← nat.prod_factors hn₀, list_prod_right hf₀] at h, + obtain ⟨p, hmem, hj⟩ := list.mem_map.mp (list.neg_one_mem_of_prod_eq_neg_one h), + exact ⟨p, nat.prime_of_mem_factors hmem, nat.dvd_of_mem_factors hmem, hj⟩, +end + end jacobi_sym namespace zmod diff --git a/src/number_theory/legendre_symbol/quadratic_reciprocity.lean b/src/number_theory/legendre_symbol/quadratic_reciprocity.lean index af92628a7b20c..946ff967ed92d 100644 --- a/src/number_theory/legendre_symbol/quadratic_reciprocity.lean +++ b/src/number_theory/legendre_symbol/quadratic_reciprocity.lean @@ -199,6 +199,71 @@ end legendre_sym end legendre +section quadratic_form + +/-! +### Applications to binary quadratic forms +-/ + +namespace legendre_sym + +/-- The Legendre symbol `legendre_sym p a = 1` if there is a solution in `ℤ/pℤ` +of the equation `x^2 - a*y^2 = 0` with `y ≠ 0`. -/ +lemma eq_one_of_sq_sub_mul_sq_eq_zero {p : ℕ} [fact p.prime] + {a : ℤ} (ha : (a : zmod p) ≠ 0) {x y : zmod p} (hy : y ≠ 0) (hxy : x ^ 2 - a * y ^ 2 = 0) : + legendre_sym p a = 1 := +begin + apply_fun (* y⁻¹ ^ 2) at hxy, + simp only [zero_mul] at hxy, + rw [(by ring : (x ^ 2 - ↑a * y ^ 2) * y⁻¹ ^ 2 = (x * y⁻¹) ^ 2 - a * (y * y⁻¹) ^ 2), + mul_inv_cancel hy, one_pow, mul_one, sub_eq_zero, pow_two] at hxy, + exact (eq_one_iff p ha).mpr ⟨x * y⁻¹, hxy.symm⟩, +end + +/-- The Legendre symbol `legendre_sym p a = 1` if there is a solution in `ℤ/pℤ` +of the equation `x^2 - a*y^2 = 0` with `x ≠ 0`. -/ +lemma eq_one_of_sq_sub_mul_sq_eq_zero' {p : ℕ} [fact p.prime] + {a : ℤ} (ha : (a : zmod p) ≠ 0) {x y : zmod p} (hx : x ≠ 0) (hxy : x ^ 2 - a * y ^ 2 = 0) : + legendre_sym p a = 1 := +begin + have hy : y ≠ 0, + { rintro rfl, + rw [zero_pow' 2 (by norm_num), mul_zero, sub_zero, pow_eq_zero_iff (by norm_num : 0 < 2)] + at hxy, + exacts [hx hxy, infer_instance], }, -- why is the instance not inferred automatically? + exact eq_one_of_sq_sub_mul_sq_eq_zero ha hy hxy, +end + +/-- If `legendre_sym p a = -1`, then the only solution of `x^2 - a*y^2 = 0` in `ℤ/pℤ` +is the trivial one. -/ +lemma eq_zero_mod_of_eq_neg_one {p : ℕ} [fact p.prime] {a : ℤ} + (h : legendre_sym p a = -1) {x y : zmod p} (hxy : x ^ 2 - a * y ^ 2 = 0) : x = 0 ∧ y = 0 := +begin + have ha : (a : zmod p) ≠ 0, + { intro hf, + rw (eq_zero_iff p a).mpr hf at h, + exact int.zero_ne_neg_of_ne zero_ne_one h, }, + by_contra hf, + cases not_and_distrib.mp hf with hx hy, + { rw [eq_one_of_sq_sub_mul_sq_eq_zero' ha hx hxy, eq_neg_self_iff] at h, + exact one_ne_zero h, }, + { rw [eq_one_of_sq_sub_mul_sq_eq_zero ha hy hxy, eq_neg_self_iff] at h, + exact one_ne_zero h, } +end + +/-- If `legendre_sym p a = -1` and `p` divides `x^2 - a*y^2`, then `p` must divide `x` and `y`. -/ +lemma prime_dvd_of_eq_neg_one {p : ℕ} [fact p.prime] {a : ℤ} + (h : legendre_sym p a = -1) {x y : ℤ} (hxy : ↑p ∣ x ^ 2 - a * y ^ 2) : ↑p ∣ x ∧ ↑p ∣ y := +begin + simp_rw ← zmod.int_coe_zmod_eq_zero_iff_dvd at hxy ⊢, + push_cast at hxy, + exact eq_zero_mod_of_eq_neg_one h hxy, +end + +end legendre_sym + +end quadratic_form + section values /-! From 4b884eff88bf72d32496b0bc7cc0bfccce77f2fa Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 15 May 2023 07:08:24 +0000 Subject: [PATCH 1010/1291] chore(measure_theory/measure/haar_of_basis): lemmas about `basis.parallelepiped` (#18873) These are bundled versions of the lemmas about `parallelepiped`. This is working towards the result that the measure on `euclidean_space` and the corresponding pi type agree. --- src/measure_theory/measure/haar_of_basis.lean | 19 ++++++++++++++++++- 1 file changed, 18 insertions(+), 1 deletion(-) diff --git a/src/measure_theory/measure/haar_of_basis.lean b/src/measure_theory/measure/haar_of_basis.lean index dfbf34e1b7ff7..3da1e80196ca2 100644 --- a/src/measure_theory/measure/haar_of_basis.lean +++ b/src/measure_theory/measure/haar_of_basis.lean @@ -179,7 +179,7 @@ end add_comm_group section normed_space -variables [normed_add_comm_group E] [normed_space ℝ E] +variables [normed_add_comm_group E] [normed_add_comm_group F] [normed_space ℝ E] [normed_space ℝ F] /-- The parallelepiped spanned by a basis, as a compact set with nonempty interior. -/ def basis.parallelepiped (b : basis ι ℝ E) : positive_compacts E := @@ -200,6 +200,23 @@ def basis.parallelepiped (b : basis ι ℝ E) : positive_compacts E := rwa [← homeomorph.image_interior, nonempty_image_iff], end } +@[simp] lemma basis.coe_parallelepiped (b : basis ι ℝ E) : + (b.parallelepiped : set E) = parallelepiped b := +rfl + +@[simp] lemma basis.parallelepiped_reindex (b : basis ι ℝ E) (e : ι ≃ ι') : + (b.reindex e).parallelepiped = b.parallelepiped := +positive_compacts.ext $ + (congr_arg parallelepiped (b.coe_reindex _)).trans (parallelepiped_comp_equiv b e.symm) + +lemma basis.parallelepiped_map (b : basis ι ℝ E) (e : E ≃ₗ[ℝ] F) : + (b.map e).parallelepiped = b.parallelepiped.map e + (by haveI := finite_dimensional.of_fintype_basis b; exact + e.to_linear_map.continuous_of_finite_dimensional) + (by haveI := finite_dimensional.of_fintype_basis (b.map e); exact + e.to_linear_map.is_open_map_of_finite_dimensional e.surjective) := +positive_compacts.ext (image_parallelepiped e.to_linear_map _).symm + variables [measurable_space E] [borel_space E] /-- The Lebesgue measure associated to a basis, giving measure `1` to the parallelepiped spanned From e9f2a838ee9090764d63f65168bb11d6ac732145 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 15 May 2023 07:08:25 +0000 Subject: [PATCH 1011/1291] chore(analysis/normed_space/pi_Lp): add `pi_Lp.continuous_linear_equiv` (#19014) This bundling is often useful to have. We keep around `pi_Lp.linear_equiv` as this is shorter to write than `pi_Lp.continuous_linear_equiv.to_linear_equiv`. --- src/analysis/normed_space/pi_Lp.lean | 9 ++++++++- 1 file changed, 8 insertions(+), 1 deletion(-) diff --git a/src/analysis/normed_space/pi_Lp.lean b/src/analysis/normed_space/pi_Lp.lean index da14f060696de..d6ead6efaefb7 100644 --- a/src/analysis/normed_space/pi_Lp.lean +++ b/src/analysis/normed_space/pi_Lp.lean @@ -707,13 +707,20 @@ lemma norm_equiv_symm_one {β} [seminormed_add_comm_group β] (hp : p ≠ ∞) [ variables (𝕜 p) -/-- `pi_Lp.equiv` as a linear map. -/ +/-- `pi_Lp.equiv` as a linear equivalence. -/ @[simps {fully_applied := ff}] protected def linear_equiv : pi_Lp p β ≃ₗ[𝕜] Π i, β i := { to_fun := pi_Lp.equiv _ _, inv_fun := (pi_Lp.equiv _ _).symm, ..linear_equiv.refl _ _} +/-- `pi_Lp.equiv` as a continuous linear equivalence. -/ +@[simps {fully_applied := ff}] +protected def continuous_linear_equiv : pi_Lp p β ≃L[𝕜] Π i, β i := +{ to_linear_equiv := pi_Lp.linear_equiv _ _ _, + continuous_to_fun := continuous_equiv _ _, + continuous_inv_fun := continuous_equiv_symm _ _ } + section basis variables (ι) From dbdf71cee7bb20367cb7e37279c08b0c218cf967 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Mon, 15 May 2023 08:40:11 +0000 Subject: [PATCH 1012/1291] chore(*): add mathlib4 synchronization comments (#19009) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.category.Group.injective` * `algebra.category.Ring.basic` * `algebra.homology.augment` * `analysis.convex.contractible` * `analysis.convex.gauge` * `category_theory.abelian.homology` * `category_theory.subterminal` * `model_theory.bundled` * `model_theory.elementary_maps` * `model_theory.finitely_generated` * `model_theory.skolem` * `topology.category.Top.limits.cofiltered` * `topology.category.Top.limits.konig` * `topology.homotopy.contractible` * `topology.homotopy.path` * `topology.sheaves.abelian` --- src/algebra/category/Group/injective.lean | 3 +++ src/algebra/category/Ring/basic.lean | 3 +++ src/algebra/homology/augment.lean | 3 +++ src/analysis/convex/contractible.lean | 3 +++ src/analysis/convex/gauge.lean | 3 +++ src/category_theory/abelian/homology.lean | 3 +++ src/category_theory/subterminal.lean | 3 +++ src/model_theory/bundled.lean | 3 +++ src/model_theory/elementary_maps.lean | 3 +++ src/model_theory/finitely_generated.lean | 3 +++ src/model_theory/skolem.lean | 3 +++ src/topology/category/Top/limits/cofiltered.lean | 3 +++ src/topology/category/Top/limits/konig.lean | 3 +++ src/topology/homotopy/contractible.lean | 3 +++ src/topology/homotopy/path.lean | 3 +++ src/topology/sheaves/abelian.lean | 3 +++ 16 files changed, 48 insertions(+) diff --git a/src/algebra/category/Group/injective.lean b/src/algebra/category/Group/injective.lean index 41b77b856a682..730283712ca55 100644 --- a/src/algebra/category/Group/injective.lean +++ b/src/algebra/category/Group/injective.lean @@ -13,6 +13,9 @@ import ring_theory.principal_ideal_domain /-! # Injective objects in the category of abelian groups +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove that divisible groups are injective object in category of (additive) abelian groups. diff --git a/src/algebra/category/Ring/basic.lean b/src/algebra/category/Ring/basic.lean index fe15d3d8ab34b..3853efd1e88ad 100644 --- a/src/algebra/category/Ring/basic.lean +++ b/src/algebra/category/Ring/basic.lean @@ -11,6 +11,9 @@ import algebra.ring.equiv /-! # Category instances for semiring, ring, comm_semiring, and comm_ring. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We introduce the bundled categories: * `SemiRing` * `Ring` diff --git a/src/algebra/homology/augment.lean b/src/algebra/homology/augment.lean index 6e560d40863e6..0eed5748bf070 100644 --- a/src/algebra/homology/augment.lean +++ b/src/algebra/homology/augment.lean @@ -7,6 +7,9 @@ import algebra.homology.single /-! # Augmentation and truncation of `ℕ`-indexed (co)chain complexes. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ noncomputable theory diff --git a/src/analysis/convex/contractible.lean b/src/analysis/convex/contractible.lean index 6d06aa1c54b9e..88bf959d51e58 100644 --- a/src/analysis/convex/contractible.lean +++ b/src/analysis/convex/contractible.lean @@ -9,6 +9,9 @@ import topology.homotopy.contractible /-! # A convex set is contractible +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove that a (star) convex set in a real topological vector space is a contractible topological space. -/ diff --git a/src/analysis/convex/gauge.lean b/src/analysis/convex/gauge.lean index ad9c43b0664dd..e027faa1a0d3d 100644 --- a/src/analysis/convex/gauge.lean +++ b/src/analysis/convex/gauge.lean @@ -12,6 +12,9 @@ import tactic.congrm /-! # The Minkowksi functional +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the Minkowski functional, aka gauge. The Minkowski functional of a set `s` is the function which associates each point to how much you diff --git a/src/category_theory/abelian/homology.lean b/src/category_theory/abelian/homology.lean index ba961216716d4..4aea6cb3442c6 100644 --- a/src/category_theory/abelian/homology.lean +++ b/src/category_theory/abelian/homology.lean @@ -9,6 +9,9 @@ import category_theory.limits.preserves.shapes.kernels import category_theory.limits.preserves.shapes.images /-! +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The object `homology f g w`, where `w : f ≫ g = 0`, can be identified with either a cokernel or a kernel. The isomorphism with a cokernel is `homology_iso_cokernel_lift`, which diff --git a/src/category_theory/subterminal.lean b/src/category_theory/subterminal.lean index 5adb0eedb5aae..0d523650bed41 100644 --- a/src/category_theory/subterminal.lean +++ b/src/category_theory/subterminal.lean @@ -10,6 +10,9 @@ import category_theory.subobject.mono_over /-! # Subterminal objects +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Subterminal objects are the objects which can be thought of as subobjects of the terminal object. In fact, the definition can be constructed to not require a terminal object, by defining `A` to be subterminal iff for any `Z`, there is at most one morphism `Z ⟶ A`. diff --git a/src/model_theory/bundled.lean b/src/model_theory/bundled.lean index bcbbb232f22c6..4f7f34dc38f06 100644 --- a/src/model_theory/bundled.lean +++ b/src/model_theory/bundled.lean @@ -7,6 +7,9 @@ import model_theory.elementary_maps import category_theory.concrete_category.bundled /-! # Bundled First-Order Structures + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. This file bundles types together with their first-order structure. ## Main Definitions diff --git a/src/model_theory/elementary_maps.lean b/src/model_theory/elementary_maps.lean index 3db612f895343..4d014c5ae8aa8 100644 --- a/src/model_theory/elementary_maps.lean +++ b/src/model_theory/elementary_maps.lean @@ -9,6 +9,9 @@ import model_theory.substructures /-! # Elementary Maps Between First-Order Structures +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main Definitions * A `first_order.language.elementary_embedding` is an embedding that commutes with the realizations of formulas. diff --git a/src/model_theory/finitely_generated.lean b/src/model_theory/finitely_generated.lean index e178102185101..2eee09bdd08a2 100644 --- a/src/model_theory/finitely_generated.lean +++ b/src/model_theory/finitely_generated.lean @@ -7,6 +7,9 @@ import model_theory.substructures /-! # Finitely Generated First-Order Structures + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. This file defines what it means for a first-order (sub)structure to be finitely or countably generated, similarly to other finitely-generated objects in the algebra library. diff --git a/src/model_theory/skolem.lean b/src/model_theory/skolem.lean index 33d7de5234f3e..517db8987361f 100644 --- a/src/model_theory/skolem.lean +++ b/src/model_theory/skolem.lean @@ -8,6 +8,9 @@ import model_theory.elementary_maps /-! # Skolem Functions and Downward Löwenheim–Skolem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main Definitions * `first_order.language.skolem₁` is a language consisting of Skolem functions for another language. diff --git a/src/topology/category/Top/limits/cofiltered.lean b/src/topology/category/Top/limits/cofiltered.lean index 980f9f23f4fe4..549bfcbcc81a3 100644 --- a/src/topology/category/Top/limits/cofiltered.lean +++ b/src/topology/category/Top/limits/cofiltered.lean @@ -8,6 +8,9 @@ import topology.category.Top.limits.basic /-! # Cofiltered limits in Top. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given a *compatible* collection of topological bases for the factors in a cofiltered limit which contain `set.univ` and are closed under intersections, the induced *naive* collection of sets in the limit is, in fact, a topological basis. diff --git a/src/topology/category/Top/limits/konig.lean b/src/topology/category/Top/limits/konig.lean index 0548dc673a9b2..992e9c8e6ed50 100644 --- a/src/topology/category/Top/limits/konig.lean +++ b/src/topology/category/Top/limits/konig.lean @@ -8,6 +8,9 @@ import topology.category.Top.limits.basic /-! # Topological Kőnig's lemma +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A topological version of Kőnig's lemma is that the inverse limit of nonempty compact Hausdorff spaces is nonempty. (Note: this can be generalized further to inverse limits of nonempty compact T0 spaces, where all the maps are closed maps; see [Stone1979] --- however there is an erratum diff --git a/src/topology/homotopy/contractible.lean b/src/topology/homotopy/contractible.lean index 77f73fb7ef846..27eac82476317 100644 --- a/src/topology/homotopy/contractible.lean +++ b/src/topology/homotopy/contractible.lean @@ -10,6 +10,9 @@ import topology.homotopy.equiv /-! # Contractible spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we define `contractible_space`, a space that is homotopy equivalent to `unit`. -/ diff --git a/src/topology/homotopy/path.lean b/src/topology/homotopy/path.lean index 9e0a73787fb85..438d71e2d45d3 100644 --- a/src/topology/homotopy/path.lean +++ b/src/topology/homotopy/path.lean @@ -11,6 +11,9 @@ import analysis.convex.basic /-! # Homotopy between paths +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we define a `homotopy` between two `path`s. In addition, we define a relation `homotopic` on `path`s, and prove that it is an equivalence relation. diff --git a/src/topology/sheaves/abelian.lean b/src/topology/sheaves/abelian.lean index 26296209a65de..c63e2191cf07f 100644 --- a/src/topology/sheaves/abelian.lean +++ b/src/topology/sheaves/abelian.lean @@ -11,6 +11,9 @@ import category_theory.sites.left_exact /-! # Category of sheaves is abelian + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. Let `C, D` be categories and `J` be a grothendieck topology on `C`, when `D` is abelian and sheafification is possible in `C`, `Sheaf J D` is abelian as well (`Sheaf_is_abelian`). From 1fd85189f32d02924c7c6aec50838953e6b797c9 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 15 May 2023 10:56:56 +0000 Subject: [PATCH 1013/1291] chore(measure_theory/measure/haar_lebesgue): remove unecessary parens (#19017) Some cleanup extracted from a larger refactor. --- src/measure_theory/measure/haar_lebesgue.lean | 34 +++++++++---------- 1 file changed, 17 insertions(+), 17 deletions(-) diff --git a/src/measure_theory/measure/haar_lebesgue.lean b/src/measure_theory/measure/haar_lebesgue.lean index 965c33384fa36..2e82c6dc5df52 100644 --- a/src/measure_theory/measure/haar_lebesgue.lean +++ b/src/measure_theory/measure/haar_lebesgue.lean @@ -322,7 +322,7 @@ equal to `μ s` times the absolute value of the determinant of `f`. -/ -/ lemma map_add_haar_smul {r : ℝ} (hr : r ≠ 0) : - measure.map ((•) r) μ = ennreal.of_real (abs (r ^ (finrank ℝ E))⁻¹) • μ := + measure.map ((•) r) μ = ennreal.of_real (abs (r ^ finrank ℝ E)⁻¹) • μ := begin let f : E →ₗ[ℝ] E := r • 1, change measure.map f μ = _, @@ -335,14 +335,14 @@ begin end @[simp] lemma add_haar_preimage_smul {r : ℝ} (hr : r ≠ 0) (s : set E) : - μ (((•) r) ⁻¹' s) = ennreal.of_real (abs (r ^ (finrank ℝ E))⁻¹) * μ s := + μ (((•) r) ⁻¹' s) = ennreal.of_real (abs (r ^ finrank ℝ E)⁻¹) * μ s := calc μ (((•) r) ⁻¹' s) = measure.map ((•) r) μ s : ((homeomorph.smul (is_unit_iff_ne_zero.2 hr).unit).to_measurable_equiv.map_apply s).symm -... = ennreal.of_real (abs (r^(finrank ℝ E))⁻¹) * μ s : by { rw map_add_haar_smul μ hr, refl } +... = ennreal.of_real (abs (r^finrank ℝ E)⁻¹) * μ s : by { rw map_add_haar_smul μ hr, refl } /-- Rescaling a set by a factor `r` multiplies its measure by `abs (r ^ dim)`. -/ @[simp] lemma add_haar_smul (r : ℝ) (s : set E) : - μ (r • s) = ennreal.of_real (abs (r ^ (finrank ℝ E))) * μ s := + μ (r • s) = ennreal.of_real (abs (r ^ finrank ℝ E)) * μ s := begin rcases ne_or_eq r 0 with h|rfl, { rw [← preimage_smul_inv₀ h, add_haar_preimage_smul μ (inv_ne_zero h), inv_pow, inv_inv] }, @@ -382,10 +382,10 @@ end variables (μ) @[simp] lemma add_haar_image_homothety (x : E) (r : ℝ) (s : set E) : - μ (affine_map.homothety x r '' s) = ennreal.of_real (abs (r ^ (finrank ℝ E))) * μ s := + μ (affine_map.homothety x r '' s) = ennreal.of_real (abs (r ^ finrank ℝ E)) * μ s := calc μ (affine_map.homothety x r '' s) = μ ((λ y, y + x) '' (r • ((λ y, y + (-x)) '' s))) : by { simp only [← image_smul, image_image, ← sub_eq_add_neg], refl } -... = ennreal.of_real (abs (r ^ (finrank ℝ E))) * μ s : +... = ennreal.of_real (abs (r ^ finrank ℝ E)) * μ s : by simp only [image_add_right, measure_preimage_add_right, add_haar_smul] /-- The integral of `f (R • x)` with respect to an additive Haar measure is a multiple of the @@ -457,7 +457,7 @@ begin end lemma add_haar_ball_mul_of_pos (x : E) {r : ℝ} (hr : 0 < r) (s : ℝ) : - μ (ball x (r * s)) = ennreal.of_real (r ^ (finrank ℝ E)) * μ (ball 0 s) := + μ (ball x (r * s)) = ennreal.of_real (r ^ finrank ℝ E) * μ (ball 0 s) := begin have : ball (0 : E) (r * s) = r • ball 0 s, by simp only [smul_ball hr.ne' (0 : E) s, real.norm_eq_abs, abs_of_nonneg hr.le, smul_zero], @@ -465,11 +465,11 @@ begin end lemma add_haar_ball_of_pos (x : E) {r : ℝ} (hr : 0 < r) : - μ (ball x r) = ennreal.of_real (r ^ (finrank ℝ E)) * μ (ball 0 1) := + μ (ball x r) = ennreal.of_real (r ^ finrank ℝ E) * μ (ball 0 1) := by rw [← add_haar_ball_mul_of_pos μ x hr, mul_one] lemma add_haar_ball_mul [nontrivial E] (x : E) {r : ℝ} (hr : 0 ≤ r) (s : ℝ) : - μ (ball x (r * s)) = ennreal.of_real (r ^ (finrank ℝ E)) * μ (ball 0 s) := + μ (ball x (r * s)) = ennreal.of_real (r ^ finrank ℝ E) * μ (ball 0 s) := begin rcases has_le.le.eq_or_lt hr with h|h, { simp only [← h, zero_pow finrank_pos, measure_empty, zero_mul, ennreal.of_real_zero, @@ -478,11 +478,11 @@ begin end lemma add_haar_ball [nontrivial E] (x : E) {r : ℝ} (hr : 0 ≤ r) : - μ (ball x r) = ennreal.of_real (r ^ (finrank ℝ E)) * μ (ball 0 1) := + μ (ball x r) = ennreal.of_real (r ^ finrank ℝ E) * μ (ball 0 1) := by rw [← add_haar_ball_mul μ x hr, mul_one] lemma add_haar_closed_ball_mul_of_pos (x : E) {r : ℝ} (hr : 0 < r) (s : ℝ) : - μ (closed_ball x (r * s)) = ennreal.of_real (r ^ (finrank ℝ E)) * μ (closed_ball 0 s) := + μ (closed_ball x (r * s)) = ennreal.of_real (r ^ finrank ℝ E) * μ (closed_ball 0 s) := begin have : closed_ball (0 : E) (r * s) = r • closed_ball 0 s, by simp [smul_closed_ball' hr.ne' (0 : E), abs_of_nonneg hr.le], @@ -490,7 +490,7 @@ begin end lemma add_haar_closed_ball_mul (x : E) {r : ℝ} (hr : 0 ≤ r) {s : ℝ} (hs : 0 ≤ s) : - μ (closed_ball x (r * s)) = ennreal.of_real (r ^ (finrank ℝ E)) * μ (closed_ball 0 s) := + μ (closed_ball x (r * s)) = ennreal.of_real (r ^ finrank ℝ E) * μ (closed_ball 0 s) := begin have : closed_ball (0 : E) (r * s) = r • closed_ball 0 s, by simp [smul_closed_ball r (0 : E) hs, abs_of_nonneg hr], @@ -501,15 +501,15 @@ end Use instead `add_haar_closed_ball`, which uses the measure of the open unit ball as a standard form. -/ lemma add_haar_closed_ball' (x : E) {r : ℝ} (hr : 0 ≤ r) : - μ (closed_ball x r) = ennreal.of_real (r ^ (finrank ℝ E)) * μ (closed_ball 0 1) := + μ (closed_ball x r) = ennreal.of_real (r ^ finrank ℝ E) * μ (closed_ball 0 1) := by rw [← add_haar_closed_ball_mul μ x hr zero_le_one, mul_one] lemma add_haar_closed_unit_ball_eq_add_haar_unit_ball : μ (closed_ball (0 : E) 1) = μ (ball 0 1) := begin apply le_antisymm _ (measure_mono ball_subset_closed_ball), - have A : tendsto (λ (r : ℝ), ennreal.of_real (r ^ (finrank ℝ E)) * μ (closed_ball (0 : E) 1)) - (𝓝[<] 1) (𝓝 (ennreal.of_real (1 ^ (finrank ℝ E)) * μ (closed_ball (0 : E) 1))), + have A : tendsto (λ (r : ℝ), ennreal.of_real (r ^ finrank ℝ E) * μ (closed_ball (0 : E) 1)) + (𝓝[<] 1) (𝓝 (ennreal.of_real (1 ^ finrank ℝ E) * μ (closed_ball (0 : E) 1))), { refine ennreal.tendsto.mul _ (by simp) tendsto_const_nhds (by simp), exact ennreal.tendsto_of_real ((tendsto_id'.2 nhds_within_le_nhds).pow _) }, simp only [one_pow, one_mul, ennreal.of_real_one] at A, @@ -521,7 +521,7 @@ begin end lemma add_haar_closed_ball (x : E) {r : ℝ} (hr : 0 ≤ r) : - μ (closed_ball x r) = ennreal.of_real (r ^ (finrank ℝ E)) * μ (ball 0 1) := + μ (closed_ball x r) = ennreal.of_real (r ^ finrank ℝ E) * μ (ball 0 1) := by rw [add_haar_closed_ball' μ x hr, add_haar_closed_unit_ball_eq_add_haar_unit_ball] lemma add_haar_closed_ball_eq_add_haar_ball [nontrivial E] (x : E) (r : ℝ) : @@ -578,7 +578,7 @@ calc @[priority 100] instance is_unif_loc_doubling_measure_of_is_add_haar_measure : is_unif_loc_doubling_measure μ := begin - refine ⟨⟨(2 : ℝ≥0) ^ (finrank ℝ E), _⟩⟩, + refine ⟨⟨(2 : ℝ≥0) ^ finrank ℝ E, _⟩⟩, filter_upwards [self_mem_nhds_within] with r hr x, rw [add_haar_closed_ball_mul μ x zero_le_two (le_of_lt hr), add_haar_closed_ball_center μ x, ennreal.of_real, real.to_nnreal_pow zero_le_two], From 24e0c85412ff6adbeca08022c25ba4876eedf37a Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Mon, 15 May 2023 16:30:23 +0000 Subject: [PATCH 1014/1291] chore(measure_theory/{integral/set_integral,function/*}): split out inner product space material (#19019) Split out inner product space material from ```text measure_theory/function/lp_space measure_theory/function/l1_space measure_theory/integral/set_integral ``` into `measure_theory/integral/l2_space`, the only place it's used. This removes the dependence of the Bochner integral file on `inner_product_space`. --- src/measure_theory/function/l1_space.lean | 14 ------ src/measure_theory/function/l2_space.lean | 46 ++++++++++++++++++- src/measure_theory/function/lp_space.lean | 24 ++-------- src/measure_theory/integral/set_integral.lean | 25 +--------- 4 files changed, 50 insertions(+), 59 deletions(-) diff --git a/src/measure_theory/function/l1_space.lean b/src/measure_theory/function/l1_space.lean index 2f1c344e116b0..89b1e64eeb428 100644 --- a/src/measure_theory/function/l1_space.lean +++ b/src/measure_theory/function/l1_space.lean @@ -1013,20 +1013,6 @@ by { rw ← mem_ℒp_one_iff_integrable at hf ⊢, exact hf.im, } end is_R_or_C -section inner_product -variables {𝕜 E : Type*} -variables [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] {f : α → E} - -local notation `⟪`x`, `y`⟫` := @inner 𝕜 E _ x y - -lemma integrable.const_inner (c : E) (hf : integrable f μ) : integrable (λ x, ⟪c, f x⟫) μ := -by { rw ← mem_ℒp_one_iff_integrable at hf ⊢, exact hf.const_inner c, } - -lemma integrable.inner_const (hf : integrable f μ) (c : E) : integrable (λ x, ⟪f x, c⟫) μ := -by { rw ← mem_ℒp_one_iff_integrable at hf ⊢, exact hf.inner_const c, } - -end inner_product - section trim variables {H : Type*} [normed_add_comm_group H] {m0 : measurable_space α} {μ' : measure α} diff --git a/src/measure_theory/function/l2_space.lean b/src/measure_theory/function/l2_space.lean index ee1620be1b073..a46364500b128 100644 --- a/src/measure_theory/function/l2_space.lean +++ b/src/measure_theory/function/l2_space.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Rémy Degenne. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ -import analysis.inner_product_space.basic +import measure_theory.function.strongly_measurable.inner import measure_theory.integral.set_integral /-! # `L^2` space @@ -22,7 +22,7 @@ is also an inner product space, with inner product defined as `inner f g = ∫ a -/ noncomputable theory -open topological_space measure_theory measure_theory.Lp +open topological_space measure_theory measure_theory.Lp filter open_locale nnreal ennreal measure_theory namespace measure_theory @@ -55,6 +55,48 @@ end end +section inner_product_space + +variables {α : Type*} {m : measurable_space α} {p : ℝ≥0∞} {μ : measure α} +variables {E 𝕜 : Type*} [is_R_or_C 𝕜] [normed_add_comm_group E] [inner_product_space 𝕜 E] + +local notation `⟪`x`, `y`⟫` := @inner 𝕜 E _ x y + +lemma mem_ℒp.const_inner (c : E) {f : α → E} (hf : mem_ℒp f p μ) : + mem_ℒp (λ a, ⟪c, f a⟫) p μ := +hf.of_le_mul (ae_strongly_measurable.inner ae_strongly_measurable_const hf.1) + (eventually_of_forall (λ x, norm_inner_le_norm _ _)) + +lemma mem_ℒp.inner_const {f : α → E} (hf : mem_ℒp f p μ) (c : E) : + mem_ℒp (λ a, ⟪f a, c⟫) p μ := +hf.of_le_mul (ae_strongly_measurable.inner hf.1 ae_strongly_measurable_const) + (eventually_of_forall (λ x, by { rw mul_comm, exact norm_inner_le_norm _ _, })) + +variables {f : α → E} + +lemma integrable.const_inner (c : E) (hf : integrable f μ) : integrable (λ x, ⟪c, f x⟫) μ := +by { rw ← mem_ℒp_one_iff_integrable at hf ⊢, exact hf.const_inner c, } + +lemma integrable.inner_const (hf : integrable f μ) (c : E) : integrable (λ x, ⟪f x, c⟫) μ := +by { rw ← mem_ℒp_one_iff_integrable at hf ⊢, exact hf.inner_const c, } + +variables [complete_space E] [normed_space ℝ E] + +lemma _root_.integral_inner {f : α → E} (hf : integrable f μ) (c : E) : + ∫ x, ⟪c, f x⟫ ∂μ = ⟪c, ∫ x, f x ∂μ⟫ := +((innerSL 𝕜 c).restrict_scalars ℝ).integral_comp_comm hf + +variables (𝕜) +-- variable binder update doesn't work for lemmas which refer to `𝕜` only via the notation +local notation (name := inner_with_explicit) `⟪`x`, `y`⟫` := @inner 𝕜 E _ x y + +lemma _root_.integral_eq_zero_of_forall_integral_inner_eq_zero (f : α → E) (hf : integrable f μ) + (hf_int : ∀ (c : E), ∫ x, ⟪c, f x⟫ ∂μ = 0) : + ∫ x, f x ∂μ = 0 := +by { specialize hf_int (∫ x, f x ∂μ), rwa [integral_inner hf, inner_self_eq_zero] at hf_int } + +end inner_product_space + namespace L2 variables {α E F 𝕜 : Type*} [is_R_or_C 𝕜] [measurable_space α] {μ : measure α} diff --git a/src/measure_theory/function/lp_space.lean b/src/measure_theory/function/lp_space.lean index 0ff13d618a5e1..ac8bfcfc86d13 100644 --- a/src/measure_theory/function/lp_space.lean +++ b/src/measure_theory/function/lp_space.lean @@ -8,7 +8,6 @@ import analysis.normed.group.hom import measure_theory.function.ess_sup import measure_theory.function.ae_eq_fun import measure_theory.integral.mean_inequalities -import measure_theory.function.strongly_measurable.inner import topology.continuous_function.compact /-! @@ -1456,35 +1455,20 @@ lemma mem_ℒp.re (hf : mem_ℒp f p μ) : mem_ℒp (λ x, is_R_or_C.re (f x)) p begin have : ∀ x, ‖is_R_or_C.re (f x)‖ ≤ 1 * ‖f x‖, by { intro x, rw one_mul, exact is_R_or_C.norm_re_le_norm (f x), }, - exact hf.of_le_mul hf.1.re (eventually_of_forall this), + refine hf.of_le_mul _ (eventually_of_forall this), + exact is_R_or_C.continuous_re.comp_ae_strongly_measurable hf.1, end lemma mem_ℒp.im (hf : mem_ℒp f p μ) : mem_ℒp (λ x, is_R_or_C.im (f x)) p μ := begin have : ∀ x, ‖is_R_or_C.im (f x)‖ ≤ 1 * ‖f x‖, by { intro x, rw one_mul, exact is_R_or_C.norm_im_le_norm (f x), }, - exact hf.of_le_mul hf.1.im (eventually_of_forall this), + refine hf.of_le_mul _ (eventually_of_forall this), + exact is_R_or_C.continuous_im.comp_ae_strongly_measurable hf.1, end end is_R_or_C -section inner_product -variables {E' 𝕜 : Type*} [is_R_or_C 𝕜] [normed_add_comm_group E'] [inner_product_space 𝕜 E'] - -local notation `⟪`x`, `y`⟫` := @inner 𝕜 E' _ x y - -lemma mem_ℒp.const_inner (c : E') {f : α → E'} (hf : mem_ℒp f p μ) : - mem_ℒp (λ a, ⟪c, f a⟫) p μ := -hf.of_le_mul (ae_strongly_measurable.inner ae_strongly_measurable_const hf.1) - (eventually_of_forall (λ x, norm_inner_le_norm _ _)) - -lemma mem_ℒp.inner_const {f : α → E'} (hf : mem_ℒp f p μ) (c : E') : - mem_ℒp (λ a, ⟪f a, c⟫) p μ := -hf.of_le_mul (ae_strongly_measurable.inner hf.1 ae_strongly_measurable_const) - (eventually_of_forall (λ x, by { rw mul_comm, exact norm_inner_le_norm _ _, })) - -end inner_product - section liminf variables [measurable_space E] [opens_measurable_space E] {R : ℝ≥0} diff --git a/src/measure_theory/integral/set_integral.lean b/src/measure_theory/integral/set_integral.lean index fc4be754323a5..e1754e0df4e65 100644 --- a/src/measure_theory/integral/set_integral.lean +++ b/src/measure_theory/integral/set_integral.lean @@ -48,6 +48,8 @@ but we reference them here because all theorems about set integrals are in this -/ +assert_not_exists inner_product_space + noncomputable theory open set filter topological_space measure_theory function open_locale classical topology interval big_operators filter ennreal nnreal measure_theory @@ -1171,29 +1173,6 @@ begin simp_rw [integrable_smul_const hc, hf, not_false_iff] } end -section inner - -variables {E' : Type*} -variables [normed_add_comm_group E'] [inner_product_space 𝕜 E'] -variables [complete_space E'] [normed_space ℝ E'] - -local notation `⟪`x`, `y`⟫` := @inner 𝕜 E' _ x y - -lemma integral_inner {f : α → E'} (hf : integrable f μ) (c : E') : - ∫ x, ⟪c, f x⟫ ∂μ = ⟪c, ∫ x, f x ∂μ⟫ := -((innerSL 𝕜 c).restrict_scalars ℝ).integral_comp_comm hf - -variables (𝕜) --- variable binder update doesn't work for lemmas which refer to `𝕜` only via the notation -local notation (name := inner_with_explicit) `⟪`x`, `y`⟫` := @inner 𝕜 E' _ x y - -lemma integral_eq_zero_of_forall_integral_inner_eq_zero (f : α → E') (hf : integrable f μ) - (hf_int : ∀ (c : E'), ∫ x, ⟪c, f x⟫ ∂μ = 0) : - ∫ x, f x ∂μ = 0 := -by { specialize hf_int (∫ x, f x ∂μ), rwa [integral_inner hf, inner_self_eq_zero] at hf_int } - -end inner - lemma integral_with_density_eq_integral_smul {f : α → ℝ≥0} (f_meas : measurable f) (g : α → E) : ∫ a, g a ∂(μ.with_density (λ x, f x)) = ∫ a, f a • g a ∂μ := From 0b9eaaa7686280fad8cce467f5c3c57ee6ce77f8 Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Mon, 15 May 2023 19:52:19 +0000 Subject: [PATCH 1015/1291] chore(analysis/special_functions): split up `pow.lean` (#19006) This PR splits up the monster file `analysis.special_functions.pow` into the following 6 files: - `pow_complex`: definition of `a ^ b` for complex a, b and basic properties (179 lines) - `pow_real`: definition of `a ^ b` for real a, b & lemmas about order / monotonicity (641 lines) - `pow_nnreal`: definitions for nnreal and ennreal types (693 lines) - `pow_asymptotics`: limiting behaviour at infinity (282 lines) - `pow_continuity`: continuity properties (500 lines) - `pow_tactic`: extensions to `positivity` and `norm_num` for powers (209 lines) No code has been added or removed, only moved around, and the only changes outside these files are adjustments to imports. --- archive/imo/imo2001_q2.lean | 2 +- src/algebra/order/complete_field.lean | 2 +- src/analysis/analytic/radius_liminf.lean | 2 +- src/analysis/convex/specific_functions.lean | 2 +- src/analysis/fourier/poisson_summation.lean | 1 + src/analysis/normed_space/spectrum.lean | 1 - .../normed_space/star/multiplier.lean | 2 +- src/analysis/p_series.lean | 2 +- src/analysis/schwartz_space.lean | 2 +- .../special_functions/compare_exp.lean | 2 +- .../special_functions/japanese_bracket.lean | 1 + src/analysis/special_functions/log/base.lean | 2 +- .../special_functions/log/monotone.lean | 2 +- src/analysis/special_functions/pow.lean | 2320 ----------------- .../special_functions/pow/asymptotics.lean | 282 ++ .../special_functions/pow/complex.lean | 179 ++ .../special_functions/pow/continuity.lean | 500 ++++ .../{pow_deriv.lean => pow/deriv.lean} | 2 +- .../special_functions/pow/nnreal.lean | 693 +++++ src/analysis/special_functions/pow/real.lean | 641 +++++ .../special_functions/pow/tactic.lean | 213 ++ src/analysis/specific_limits/floor_pow.lean | 2 +- src/combinatorics/additive/behrend.lean | 1 + .../simple_graph/regularity/bound.lean | 2 +- .../quadratic_form/complex.lean | 2 +- src/linear_algebra/quadratic_form/real.lean | 2 +- .../function/convergence_in_measure.lean | 2 +- .../function/special_functions/basic.lean | 2 +- src/measure_theory/measure/hausdorff.lean | 1 - src/number_theory/bertrand.lean | 1 + .../admissible_card_pow_degree.lean | 8 +- src/number_theory/class_number/finite.lean | 2 +- .../liouville/liouville_with.lean | 2 +- src/ring_theory/perfection.lean | 2 +- src/topology/metric_space/holder.lean | 2 +- test/norm_num_ext.lean | 2 +- test/positivity.lean | 2 +- 37 files changed, 2539 insertions(+), 2349 deletions(-) delete mode 100644 src/analysis/special_functions/pow.lean create mode 100644 src/analysis/special_functions/pow/asymptotics.lean create mode 100644 src/analysis/special_functions/pow/complex.lean create mode 100644 src/analysis/special_functions/pow/continuity.lean rename src/analysis/special_functions/{pow_deriv.lean => pow/deriv.lean} (99%) create mode 100644 src/analysis/special_functions/pow/nnreal.lean create mode 100644 src/analysis/special_functions/pow/real.lean create mode 100644 src/analysis/special_functions/pow/tactic.lean diff --git a/archive/imo/imo2001_q2.lean b/archive/imo/imo2001_q2.lean index 9a2ec6f7c208b..a47afd9d539ac 100644 --- a/archive/imo/imo2001_q2.lean +++ b/archive/imo/imo2001_q2.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Tian Chen -/ -import analysis.special_functions.pow +import analysis.special_functions.pow.real /-! # IMO 2001 Q2 diff --git a/src/algebra/order/complete_field.lean b/src/algebra/order/complete_field.lean index 1b886cb7392cb..5e62fce7c5440 100644 --- a/src/algebra/order/complete_field.lean +++ b/src/algebra/order/complete_field.lean @@ -5,7 +5,7 @@ Authors: Alex J. Best, Yaël Dillies -/ import algebra.order.hom.ring import algebra.order.pointwise -import analysis.special_functions.pow +import analysis.special_functions.pow.real /-! # Conditionally complete linear ordered fields diff --git a/src/analysis/analytic/radius_liminf.lean b/src/analysis/analytic/radius_liminf.lean index bd745bebaee50..949ad9c17f72e 100644 --- a/src/analysis/analytic/radius_liminf.lean +++ b/src/analysis/analytic/radius_liminf.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import analysis.analytic.basic -import analysis.special_functions.pow +import analysis.special_functions.pow.nnreal /-! # Representation of `formal_multilinear_series.radius` as a `liminf` diff --git a/src/analysis/convex/specific_functions.lean b/src/analysis/convex/specific_functions.lean index 8a07f38210f91..f6e81716af065 100644 --- a/src/analysis/convex/specific_functions.lean +++ b/src/analysis/convex/specific_functions.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov, Sébastien Gouëzel -/ -import analysis.special_functions.pow_deriv +import analysis.special_functions.pow.deriv import analysis.special_functions.sqrt /-! diff --git a/src/analysis/fourier/poisson_summation.lean b/src/analysis/fourier/poisson_summation.lean index 17e5fae7dc9cc..682068c630569 100644 --- a/src/analysis/fourier/poisson_summation.lean +++ b/src/analysis/fourier/poisson_summation.lean @@ -7,6 +7,7 @@ import analysis.fourier.add_circle import analysis.fourier.fourier_transform import analysis.p_series import analysis.schwartz_space +import analysis.special_functions.pow.tactic /-! # Poisson's summation formula diff --git a/src/analysis/normed_space/spectrum.lean b/src/analysis/normed_space/spectrum.lean index cab94d955eebb..13a1e0bf880d1 100644 --- a/src/analysis/normed_space/spectrum.lean +++ b/src/analysis/normed_space/spectrum.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jireh Loreaux -/ import algebra.algebra.spectrum -import analysis.special_functions.pow import analysis.complex.liouville import analysis.complex.polynomial import analysis.analytic.radius_liminf diff --git a/src/analysis/normed_space/star/multiplier.lean b/src/analysis/normed_space/star/multiplier.lean index 039c59151145b..26302dcbd1276 100644 --- a/src/analysis/normed_space/star/multiplier.lean +++ b/src/analysis/normed_space/star/multiplier.lean @@ -7,7 +7,7 @@ Authors: Jireh Loreaux, Jon Bannon import algebra.star.star_alg_hom import analysis.normed_space.star.basic import analysis.normed_space.operator_norm -import analysis.special_functions.pow +import analysis.special_functions.pow.nnreal import analysis.normed_space.star.mul /-! diff --git a/src/analysis/p_series.lean b/src/analysis/p_series.lean index 3418373f0dab9..ee72f8505e915 100644 --- a/src/analysis/p_series.lean +++ b/src/analysis/p_series.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 Yury G. Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury G. Kudryashov -/ -import analysis.special_functions.pow +import analysis.special_functions.pow.nnreal /-! # Convergence of `p`-series diff --git a/src/analysis/schwartz_space.lean b/src/analysis/schwartz_space.lean index ccfc4a366a59e..42b870cd55355 100644 --- a/src/analysis/schwartz_space.lean +++ b/src/analysis/schwartz_space.lean @@ -10,7 +10,7 @@ import analysis.locally_convex.with_seminorms import topology.algebra.uniform_filter_basis import topology.continuous_function.bounded import tactic.positivity -import analysis.special_functions.pow +import analysis.special_functions.pow.real /-! # Schwartz space diff --git a/src/analysis/special_functions/compare_exp.lean b/src/analysis/special_functions/compare_exp.lean index 9df5efde21bea..87fb46b1c9c32 100644 --- a/src/analysis/special_functions/compare_exp.lean +++ b/src/analysis/special_functions/compare_exp.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ -import analysis.special_functions.pow +import analysis.special_functions.pow.asymptotics import analysis.asymptotics.asymptotic_equivalent import analysis.asymptotics.specific_asymptotics diff --git a/src/analysis/special_functions/japanese_bracket.lean b/src/analysis/special_functions/japanese_bracket.lean index d287916d88f20..6b6808010278e 100644 --- a/src/analysis/special_functions/japanese_bracket.lean +++ b/src/analysis/special_functions/japanese_bracket.lean @@ -5,6 +5,7 @@ Authors: Moritz Doll -/ import measure_theory.measure.haar_lebesgue import measure_theory.integral.layercake +import analysis.special_functions.pow.tactic /-! # Japanese Bracket diff --git a/src/analysis/special_functions/log/base.lean b/src/analysis/special_functions/log/base.lean index 55ad5d89c3c36..03073d7be214a 100644 --- a/src/analysis/special_functions/log/base.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.pow +import analysis.special_functions.pow.real import data.int.log /-! diff --git a/src/analysis/special_functions/log/monotone.lean b/src/analysis/special_functions/log/monotone.lean index f0981e4f14658..9eb14b1c14252 100644 --- a/src/analysis/special_functions/log/monotone.lean +++ b/src/analysis/special_functions/log/monotone.lean @@ -3,7 +3,7 @@ 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.pow +import analysis.special_functions.pow.real /-! # Logarithm Tonality diff --git a/src/analysis/special_functions/pow.lean b/src/analysis/special_functions/pow.lean deleted file mode 100644 index 21d391547585d..0000000000000 --- a/src/analysis/special_functions/pow.lean +++ /dev/null @@ -1,2320 +0,0 @@ -/- -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, Sébastien Gouëzel, - Rémy Degenne, David Loeffler --/ -import analysis.special_functions.complex.log - -/-! -# Power function on `ℂ`, `ℝ`, `ℝ≥0`, and `ℝ≥0∞` - -We construct the power functions `x ^ y` where -* `x` and `y` are complex numbers, -* or `x` and `y` are real numbers, -* or `x` is a nonnegative real number and `y` is a real number; -* or `x` is a number from `[0, +∞]` (a.k.a. `ℝ≥0∞`) and `y` is a real number. - -We also prove basic properties of these functions. --/ - -noncomputable theory - -open_locale classical real topology nnreal ennreal filter big_operators complex_conjugate -open filter finset set - -namespace complex - -/-- The complex power function `x^y`, given by `x^y = exp(y log x)` (where `log` is the principal -determination of the logarithm), unless `x = 0` where one sets `0^0 = 1` and `0^y = 0` for -`y ≠ 0`. -/ -noncomputable def cpow (x y : ℂ) : ℂ := -if x = 0 - then if y = 0 - then 1 - else 0 - else exp (log x * y) - -noncomputable instance : has_pow ℂ ℂ := ⟨cpow⟩ - -@[simp] lemma cpow_eq_pow (x y : ℂ) : cpow x y = x ^ y := rfl - -lemma cpow_def (x y : ℂ) : x ^ y = - if x = 0 - then if y = 0 - then 1 - else 0 - else exp (log x * y) := rfl - -lemma cpow_def_of_ne_zero {x : ℂ} (hx : x ≠ 0) (y : ℂ) : x ^ y = exp (log x * y) := if_neg hx - -@[simp] lemma cpow_zero (x : ℂ) : x ^ (0 : ℂ) = 1 := by simp [cpow_def] - -@[simp] lemma cpow_eq_zero_iff (x y : ℂ) : x ^ y = 0 ↔ x = 0 ∧ y ≠ 0 := -by { simp only [cpow_def], split_ifs; simp [*, exp_ne_zero] } - -@[simp] lemma zero_cpow {x : ℂ} (h : x ≠ 0) : (0 : ℂ) ^ x = 0 := -by simp [cpow_def, *] - -lemma zero_cpow_eq_iff {x : ℂ} {a : ℂ} : 0 ^ x = a ↔ (x ≠ 0 ∧ a = 0) ∨ (x = 0 ∧ a = 1) := -begin - split, - { intros hyp, - simp only [cpow_def, eq_self_iff_true, if_true] at hyp, - by_cases x = 0, - { subst h, simp only [if_true, eq_self_iff_true] at hyp, right, exact ⟨rfl, hyp.symm⟩}, - { rw if_neg h at hyp, left, exact ⟨h, hyp.symm⟩, }, }, - { rintro (⟨h, rfl⟩|⟨rfl,rfl⟩), - { exact zero_cpow h, }, - { exact cpow_zero _, }, }, -end - -lemma eq_zero_cpow_iff {x : ℂ} {a : ℂ} : a = 0 ^ x ↔ (x ≠ 0 ∧ a = 0) ∨ (x = 0 ∧ a = 1) := -by rw [←zero_cpow_eq_iff, eq_comm] - -@[simp] lemma cpow_one (x : ℂ) : x ^ (1 : ℂ) = x := -if hx : x = 0 then by simp [hx, cpow_def] -else by rw [cpow_def, if_neg (one_ne_zero : (1 : ℂ) ≠ 0), if_neg hx, mul_one, exp_log hx] - -@[simp] lemma one_cpow (x : ℂ) : (1 : ℂ) ^ x = 1 := -by rw cpow_def; split_ifs; simp [one_ne_zero, *] at * - -lemma cpow_add {x : ℂ} (y z : ℂ) (hx : x ≠ 0) : x ^ (y + z) = x ^ y * x ^ z := -by simp only [cpow_def, ite_mul, boole_mul, mul_ite, mul_boole]; simp [*, exp_add, mul_add] at * - -lemma cpow_mul {x y : ℂ} (z : ℂ) (h₁ : -π < (log x * y).im) (h₂ : (log x * y).im ≤ π) : - x ^ (y * z) = (x ^ y) ^ z := -begin - simp only [cpow_def], - split_ifs; - simp [*, exp_ne_zero, log_exp h₁ h₂, mul_assoc] at * -end - -lemma cpow_neg (x y : ℂ) : x ^ -y = (x ^ y)⁻¹ := -by simp only [cpow_def, neg_eq_zero, mul_neg]; split_ifs; simp [exp_neg] - -lemma cpow_sub {x : ℂ} (y z : ℂ) (hx : x ≠ 0) : x ^ (y - z) = x ^ y / x ^ z := -by rw [sub_eq_add_neg, cpow_add _ _ hx, cpow_neg, div_eq_mul_inv] - -lemma cpow_neg_one (x : ℂ) : x ^ (-1 : ℂ) = x⁻¹ := -by simpa using cpow_neg x 1 - -@[simp, norm_cast] lemma cpow_nat_cast (x : ℂ) : ∀ (n : ℕ), x ^ (n : ℂ) = x ^ n -| 0 := by simp -| (n + 1) := if hx : x = 0 then by simp only [hx, pow_succ, - complex.zero_cpow (nat.cast_ne_zero.2 (nat.succ_ne_zero _)), zero_mul] - else by simp [cpow_add, hx, pow_add, cpow_nat_cast n] - -@[simp] lemma cpow_two (x : ℂ) : x ^ (2 : ℂ) = x ^ 2 := -by { rw ← cpow_nat_cast, simp only [nat.cast_bit0, nat.cast_one] } - -@[simp, norm_cast] lemma cpow_int_cast (x : ℂ) : ∀ (n : ℤ), x ^ (n : ℂ) = x ^ n -| (n : ℕ) := by simp -| -[1+ n] := by rw zpow_neg_succ_of_nat; - simp only [int.neg_succ_of_nat_coe, int.cast_neg, complex.cpow_neg, inv_eq_one_div, - int.cast_coe_nat, cpow_nat_cast] - -lemma cpow_nat_inv_pow (x : ℂ) {n : ℕ} (hn : n ≠ 0) : (x ^ (n⁻¹ : ℂ)) ^ n = x := -begin - suffices : im (log x * n⁻¹) ∈ Ioc (-π) π, - { rw [← cpow_nat_cast, ← cpow_mul _ this.1 this.2, inv_mul_cancel, cpow_one], - exact_mod_cast hn }, - rw [mul_comm, ← of_real_nat_cast, ← of_real_inv, of_real_mul_im, ← div_eq_inv_mul], - rw [← pos_iff_ne_zero] at hn, - have hn' : 0 < (n : ℝ), by assumption_mod_cast, - have hn1 : 1 ≤ (n : ℝ), by exact_mod_cast (nat.succ_le_iff.2 hn), - split, - { rw lt_div_iff hn', - calc -π * n ≤ -π * 1 : mul_le_mul_of_nonpos_left hn1 (neg_nonpos.2 real.pi_pos.le) - ... = -π : mul_one _ - ... < im (log x) : neg_pi_lt_log_im _ }, - { rw div_le_iff hn', - calc im (log x) ≤ π : log_im_le_pi _ - ... = π * 1 : (mul_one π).symm - ... ≤ π * n : mul_le_mul_of_nonneg_left hn1 real.pi_pos.le } -end - -lemma mul_cpow_of_real_nonneg {a b : ℝ} (ha : 0 ≤ a) (hb : 0 ≤ b) (r : ℂ) : - ((a : ℂ) * (b : ℂ)) ^ r = (a : ℂ) ^ r * (b : ℂ) ^ r := -begin - rcases eq_or_ne r 0 with rfl | hr, - { simp only [cpow_zero, mul_one] }, - rcases eq_or_lt_of_le ha with rfl | ha', - { rw [of_real_zero, zero_mul, zero_cpow hr, zero_mul] }, - rcases eq_or_lt_of_le hb with rfl | hb', - { rw [of_real_zero, mul_zero, zero_cpow hr, mul_zero] }, - have ha'' : (a : ℂ) ≠ 0 := of_real_ne_zero.mpr ha'.ne', - have hb'' : (b : ℂ) ≠ 0 := of_real_ne_zero.mpr hb'.ne', - rw [cpow_def_of_ne_zero (mul_ne_zero ha'' hb''), log_of_real_mul ha' hb'', of_real_log ha, - add_mul, exp_add, ←cpow_def_of_ne_zero ha'', ←cpow_def_of_ne_zero hb''] -end - -end complex - -section lim - -open complex - -variables {α : Type*} - -lemma zero_cpow_eq_nhds {b : ℂ} (hb : b ≠ 0) : - (λ (x : ℂ), (0 : ℂ) ^ x) =ᶠ[𝓝 b] 0 := -begin - suffices : ∀ᶠ (x : ℂ) in (𝓝 b), x ≠ 0, - from this.mono (λ x hx, by { dsimp only, rw [zero_cpow hx, pi.zero_apply]} ), - exact is_open.eventually_mem is_open_ne hb, -end - -lemma cpow_eq_nhds {a b : ℂ} (ha : a ≠ 0) : - (λ x, x ^ b) =ᶠ[𝓝 a] λ x, exp (log x * b) := -begin - suffices : ∀ᶠ (x : ℂ) in (𝓝 a), x ≠ 0, - from this.mono (λ x hx, by { dsimp only, rw [cpow_def_of_ne_zero hx], }), - exact is_open.eventually_mem is_open_ne ha, -end - -lemma cpow_eq_nhds' {p : ℂ × ℂ} (hp_fst : p.fst ≠ 0) : - (λ x, x.1 ^ x.2) =ᶠ[𝓝 p] λ x, exp (log x.1 * x.2) := -begin - suffices : ∀ᶠ (x : ℂ × ℂ) in (𝓝 p), x.1 ≠ 0, - from this.mono (λ x hx, by { dsimp only, rw cpow_def_of_ne_zero hx, }), - refine is_open.eventually_mem _ hp_fst, - change is_open {x : ℂ × ℂ | x.1 = 0}ᶜ, - rw is_open_compl_iff, - exact is_closed_eq continuous_fst continuous_const, -end - -/- Continuity of `λ x, a ^ x`: union of these two lemmas is optimal. -/ - -lemma continuous_at_const_cpow {a b : ℂ} (ha : a ≠ 0) : continuous_at (λ x, a ^ x) b := -begin - have cpow_eq : (λ x:ℂ, a ^ x) = λ x, exp (log a * x), - by { ext1 b, rw [cpow_def_of_ne_zero ha], }, - rw cpow_eq, - exact continuous_exp.continuous_at.comp (continuous_at.mul continuous_at_const continuous_at_id), -end - -lemma continuous_at_const_cpow' {a b : ℂ} (h : b ≠ 0) : continuous_at (λ x, a ^ x) b := -begin - by_cases ha : a = 0, - { rw [ha, continuous_at_congr (zero_cpow_eq_nhds h)], exact continuous_at_const, }, - { exact continuous_at_const_cpow ha, }, -end - -/-- The function `z ^ w` is continuous in `(z, w)` provided that `z` does not belong to the interval -`(-∞, 0]` on the real line. See also `complex.continuous_at_cpow_zero_of_re_pos` for a version that -works for `z = 0` but assumes `0 < re w`. -/ -lemma continuous_at_cpow {p : ℂ × ℂ} (hp_fst : 0 < p.fst.re ∨ p.fst.im ≠ 0) : - continuous_at (λ x : ℂ × ℂ, x.1 ^ x.2) p := -begin - have hp_fst_ne_zero : p.fst ≠ 0, - by { intro h, cases hp_fst; { rw h at hp_fst, simpa using hp_fst, }, }, - rw continuous_at_congr (cpow_eq_nhds' hp_fst_ne_zero), - refine continuous_exp.continuous_at.comp _, - refine continuous_at.mul (continuous_at.comp _ continuous_fst.continuous_at) - continuous_snd.continuous_at, - exact continuous_at_clog hp_fst, -end - -lemma continuous_at_cpow_const {a b : ℂ} (ha : 0 < a.re ∨ a.im ≠ 0) : - continuous_at (λ x, cpow x b) a := -tendsto.comp (@continuous_at_cpow (a, b) ha) (continuous_at_id.prod continuous_at_const) - -lemma filter.tendsto.cpow {l : filter α} {f g : α → ℂ} {a b : ℂ} (hf : tendsto f l (𝓝 a)) - (hg : tendsto g l (𝓝 b)) (ha : 0 < a.re ∨ a.im ≠ 0) : - tendsto (λ x, f x ^ g x) l (𝓝 (a ^ b)) := -(@continuous_at_cpow (a,b) ha).tendsto.comp (hf.prod_mk_nhds hg) - -lemma filter.tendsto.const_cpow {l : filter α} {f : α → ℂ} {a b : ℂ} (hf : tendsto f l (𝓝 b)) - (h : a ≠ 0 ∨ b ≠ 0) : - tendsto (λ x, a ^ f x) l (𝓝 (a ^ b)) := -begin - cases h, - { exact (continuous_at_const_cpow h).tendsto.comp hf, }, - { exact (continuous_at_const_cpow' h).tendsto.comp hf, }, -end - -variables [topological_space α] {f g : α → ℂ} {s : set α} {a : α} - -lemma continuous_within_at.cpow (hf : continuous_within_at f s a) (hg : continuous_within_at g s a) - (h0 : 0 < (f a).re ∨ (f a).im ≠ 0) : - continuous_within_at (λ x, f x ^ g x) s a := -hf.cpow hg h0 - -lemma continuous_within_at.const_cpow {b : ℂ} (hf : continuous_within_at f s a) - (h : b ≠ 0 ∨ f a ≠ 0) : - continuous_within_at (λ x, b ^ f x) s a := -hf.const_cpow h - -lemma continuous_at.cpow (hf : continuous_at f a) (hg : continuous_at g a) - (h0 : 0 < (f a).re ∨ (f a).im ≠ 0) : - continuous_at (λ x, f x ^ g x) a := -hf.cpow hg h0 - -lemma continuous_at.const_cpow {b : ℂ} (hf : continuous_at f a) (h : b ≠ 0 ∨ f a ≠ 0) : - continuous_at (λ x, b ^ f x) a := -hf.const_cpow h - -lemma continuous_on.cpow (hf : continuous_on f s) (hg : continuous_on g s) - (h0 : ∀ a ∈ s, 0 < (f a).re ∨ (f a).im ≠ 0) : - continuous_on (λ x, f x ^ g x) s := -λ a ha, (hf a ha).cpow (hg a ha) (h0 a ha) - -lemma continuous_on.const_cpow {b : ℂ} (hf : continuous_on f s) (h : b ≠ 0 ∨ ∀ a ∈ s, f a ≠ 0) : - continuous_on (λ x, b ^ f x) s := -λ a ha, (hf a ha).const_cpow (h.imp id $ λ h, h a ha) - -lemma continuous.cpow (hf : continuous f) (hg : continuous g) - (h0 : ∀ a, 0 < (f a).re ∨ (f a).im ≠ 0) : - continuous (λ x, f x ^ g x) := -continuous_iff_continuous_at.2 $ λ a, (hf.continuous_at.cpow hg.continuous_at (h0 a)) - -lemma continuous.const_cpow {b : ℂ} (hf : continuous f) (h : b ≠ 0 ∨ ∀ a, f a ≠ 0) : - continuous (λ x, b ^ f x) := -continuous_iff_continuous_at.2 $ λ a, (hf.continuous_at.const_cpow $ h.imp id $ λ h, h a) - -lemma continuous_on.cpow_const {b : ℂ} (hf : continuous_on f s) - (h : ∀ (a : α), a ∈ s → 0 < (f a).re ∨ (f a).im ≠ 0) : - continuous_on (λ x, (f x) ^ b) s := -hf.cpow continuous_on_const h - -end lim - -namespace real - -/-- The real power function `x^y`, defined as the real part of the complex power function. -For `x > 0`, it is equal to `exp(y log x)`. For `x = 0`, one sets `0^0=1` and `0^y=0` for `y ≠ 0`. -For `x < 0`, the definition is somewhat arbitary as it depends on the choice of a complex -determination of the logarithm. With our conventions, it is equal to `exp (y log x) cos (πy)`. -/ -noncomputable def rpow (x y : ℝ) := ((x : ℂ) ^ (y : ℂ)).re - -noncomputable instance : has_pow ℝ ℝ := ⟨rpow⟩ - -@[simp] lemma rpow_eq_pow (x y : ℝ) : rpow x y = x ^ y := rfl - -lemma rpow_def (x y : ℝ) : x ^ y = ((x : ℂ) ^ (y : ℂ)).re := rfl - -lemma rpow_def_of_nonneg {x : ℝ} (hx : 0 ≤ x) (y : ℝ) : x ^ y = - if x = 0 - then if y = 0 - then 1 - else 0 - else exp (log x * y) := -by simp only [rpow_def, complex.cpow_def]; - split_ifs; - simp [*, (complex.of_real_log hx).symm, -complex.of_real_mul, -is_R_or_C.of_real_mul, - (complex.of_real_mul _ _).symm, complex.exp_of_real_re] at * - -lemma rpow_def_of_pos {x : ℝ} (hx : 0 < x) (y : ℝ) : x ^ y = exp (log x * y) := -by rw [rpow_def_of_nonneg (le_of_lt hx), if_neg (ne_of_gt hx)] - -lemma exp_mul (x y : ℝ) : exp (x * y) = (exp x) ^ y := -by rw [rpow_def_of_pos (exp_pos _), log_exp] - -@[simp] lemma exp_one_rpow (x : ℝ) : exp 1 ^ x = exp x := by rw [←exp_mul, one_mul] - -lemma rpow_eq_zero_iff_of_nonneg {x y : ℝ} (hx : 0 ≤ x) : x ^ y = 0 ↔ x = 0 ∧ y ≠ 0 := -by { simp only [rpow_def_of_nonneg hx], split_ifs; simp [*, exp_ne_zero] } - -open_locale real - -lemma rpow_def_of_neg {x : ℝ} (hx : x < 0) (y : ℝ) : x ^ y = exp (log x * y) * cos (y * π) := -begin - rw [rpow_def, complex.cpow_def, if_neg], - have : complex.log x * y = ↑(log(-x) * y) + ↑(y * π) * complex.I, - { simp only [complex.log, abs_of_neg hx, complex.arg_of_real_of_neg hx, - complex.abs_of_real, complex.of_real_mul], ring }, - { rw [this, complex.exp_add_mul_I, ← complex.of_real_exp, ← complex.of_real_cos, - ← complex.of_real_sin, mul_add, ← complex.of_real_mul, ← mul_assoc, ← complex.of_real_mul, - complex.add_re, complex.of_real_re, complex.mul_re, complex.I_re, complex.of_real_im, - real.log_neg_eq_log], - ring }, - { rw complex.of_real_eq_zero, exact ne_of_lt hx } -end - -lemma rpow_def_of_nonpos {x : ℝ} (hx : x ≤ 0) (y : ℝ) : x ^ y = - if x = 0 - then if y = 0 - then 1 - else 0 - else exp (log x * y) * cos (y * π) := -by split_ifs; simp [rpow_def, *]; exact rpow_def_of_neg (lt_of_le_of_ne hx h) _ - -lemma rpow_pos_of_pos {x : ℝ} (hx : 0 < x) (y : ℝ) : 0 < x ^ y := -by rw rpow_def_of_pos hx; apply exp_pos - -@[simp] lemma rpow_zero (x : ℝ) : x ^ (0 : ℝ) = 1 := by simp [rpow_def] - -@[simp] lemma zero_rpow {x : ℝ} (h : x ≠ 0) : (0 : ℝ) ^ x = 0 := -by simp [rpow_def, *] - -lemma zero_rpow_eq_iff {x : ℝ} {a : ℝ} : 0 ^ x = a ↔ (x ≠ 0 ∧ a = 0) ∨ (x = 0 ∧ a = 1) := -begin - split, - { intros hyp, - simp only [rpow_def, complex.of_real_zero] at hyp, - by_cases x = 0, - { subst h, - simp only [complex.one_re, complex.of_real_zero, complex.cpow_zero] at hyp, - exact or.inr ⟨rfl, hyp.symm⟩}, - { rw complex.zero_cpow (complex.of_real_ne_zero.mpr h) at hyp, - exact or.inl ⟨h, hyp.symm⟩, }, }, - { rintro (⟨h,rfl⟩|⟨rfl,rfl⟩), - { exact zero_rpow h, }, - { exact rpow_zero _, }, }, -end - -lemma eq_zero_rpow_iff {x : ℝ} {a : ℝ} : a = 0 ^ x ↔ (x ≠ 0 ∧ a = 0) ∨ (x = 0 ∧ a = 1) := -by rw [←zero_rpow_eq_iff, eq_comm] - -@[simp] lemma rpow_one (x : ℝ) : x ^ (1 : ℝ) = x := by simp [rpow_def] - -@[simp] lemma one_rpow (x : ℝ) : (1 : ℝ) ^ x = 1 := by simp [rpow_def] - -lemma zero_rpow_le_one (x : ℝ) : (0 : ℝ) ^ x ≤ 1 := -by { by_cases h : x = 0; simp [h, zero_le_one] } - -lemma zero_rpow_nonneg (x : ℝ) : 0 ≤ (0 : ℝ) ^ x := -by { by_cases h : x = 0; simp [h, zero_le_one] } - -lemma rpow_nonneg_of_nonneg {x : ℝ} (hx : 0 ≤ x) (y : ℝ) : 0 ≤ x ^ y := -by rw [rpow_def_of_nonneg hx]; - split_ifs; simp only [zero_le_one, le_refl, le_of_lt (exp_pos _)] - -lemma abs_rpow_of_nonneg {x y : ℝ} (hx_nonneg : 0 ≤ x) : |x ^ y| = |x| ^ y := -begin - have h_rpow_nonneg : 0 ≤ x ^ y, from real.rpow_nonneg_of_nonneg hx_nonneg _, - rw [abs_eq_self.mpr hx_nonneg, abs_eq_self.mpr h_rpow_nonneg], -end - -lemma abs_rpow_le_abs_rpow (x y : ℝ) : |x ^ y| ≤ |x| ^ y := -begin - cases le_or_lt 0 x with hx hx, - { rw [abs_rpow_of_nonneg hx] }, - { rw [abs_of_neg hx, rpow_def_of_neg hx, rpow_def_of_pos (neg_pos.2 hx), log_neg_eq_log, - abs_mul, abs_of_pos (exp_pos _)], - exact mul_le_of_le_one_right (exp_pos _).le (abs_cos_le_one _) } -end - -lemma abs_rpow_le_exp_log_mul (x y : ℝ) : |x ^ y| ≤ exp (log x * y) := -begin - refine (abs_rpow_le_abs_rpow x y).trans _, - by_cases hx : x = 0, - { by_cases hy : y = 0; simp [hx, hy, zero_le_one] }, - { rw [rpow_def_of_pos (abs_pos.2 hx), log_abs] } -end - -lemma norm_rpow_of_nonneg {x y : ℝ} (hx_nonneg : 0 ≤ x) : ‖x ^ y‖ = ‖x‖ ^ y := -by { simp_rw real.norm_eq_abs, exact abs_rpow_of_nonneg hx_nonneg, } - -end real - -namespace complex - -lemma of_real_cpow {x : ℝ} (hx : 0 ≤ x) (y : ℝ) : ((x ^ y : ℝ) : ℂ) = (x : ℂ) ^ (y : ℂ) := -by simp only [real.rpow_def_of_nonneg hx, complex.cpow_def, of_real_eq_zero]; split_ifs; - simp [complex.of_real_log hx] - -lemma of_real_cpow_of_nonpos {x : ℝ} (hx : x ≤ 0) (y : ℂ) : - (x : ℂ) ^ y = ((-x) : ℂ) ^ y * exp (π * I * y) := -begin - rcases hx.eq_or_lt with rfl|hlt, - { rcases eq_or_ne y 0 with rfl|hy; simp * }, - have hne : (x : ℂ) ≠ 0, from of_real_ne_zero.mpr hlt.ne, - rw [cpow_def_of_ne_zero hne, cpow_def_of_ne_zero (neg_ne_zero.2 hne), ← exp_add, ← add_mul, - log, log, abs.map_neg, arg_of_real_of_neg hlt, ← of_real_neg, - arg_of_real_of_nonneg (neg_nonneg.2 hx), of_real_zero, zero_mul, add_zero] -end - -lemma abs_cpow_of_ne_zero {z : ℂ} (hz : z ≠ 0) (w : ℂ) : - abs (z ^ w) = abs z ^ w.re / real.exp (arg z * im w) := -by rw [cpow_def_of_ne_zero hz, abs_exp, mul_re, log_re, log_im, real.exp_sub, - real.rpow_def_of_pos (abs.pos hz)] - -lemma abs_cpow_of_imp {z w : ℂ} (h : z = 0 → w.re = 0 → w = 0) : - abs (z ^ w) = abs z ^ w.re / real.exp (arg z * im w) := -begin - rcases ne_or_eq z 0 with hz|rfl; [exact (abs_cpow_of_ne_zero hz w), rw map_zero], - cases eq_or_ne w.re 0 with hw hw, - { simp [hw, h rfl hw] }, - { rw [real.zero_rpow hw, zero_div, zero_cpow, map_zero], - exact ne_of_apply_ne re hw } -end - -lemma abs_cpow_le (z w : ℂ) : abs (z ^ w) ≤ abs z ^ w.re / real.exp (arg z * im w) := -begin - rcases ne_or_eq z 0 with hz|rfl; [exact (abs_cpow_of_ne_zero hz w).le, rw map_zero], - rcases eq_or_ne w 0 with rfl|hw, { simp }, - rw [zero_cpow hw, map_zero], - exact div_nonneg (real.rpow_nonneg_of_nonneg le_rfl _) (real.exp_pos _).le -end - -section - -variables {α : Type*} {l : filter α} {f g : α → ℂ} - -open asymptotics - -lemma is_Theta_exp_arg_mul_im (hl : is_bounded_under (≤) l (λ x, |(g x).im|)) : - (λ x, real.exp (arg (f x) * im (g x))) =Θ[l] (λ x, (1 : ℝ)) := -begin - rcases hl with ⟨b, hb⟩, - refine real.is_Theta_exp_comp_one.2 ⟨π * b, _⟩, - rw eventually_map at hb ⊢, - refine hb.mono (λ x hx, _), - erw [abs_mul], - exact mul_le_mul (abs_arg_le_pi _) hx (abs_nonneg _) real.pi_pos.le -end - -lemma is_O_cpow_rpow (hl : is_bounded_under (≤) l (λ x, |(g x).im|)) : - (λ x, f x ^ g x) =O[l] (λ x, abs (f x) ^ (g x).re) := -calc (λ x, f x ^ g x) =O[l] (λ x, abs (f x) ^ (g x).re / real.exp (arg (f x) * im (g x))) : - is_O_of_le _ $ λ x, (abs_cpow_le _ _).trans (le_abs_self _) -... =Θ[l] (λ x, abs (f x) ^ (g x).re / (1 : ℝ)) : - (is_Theta_refl _ _).div (is_Theta_exp_arg_mul_im hl) -... =ᶠ[l] (λ x, abs (f x) ^ (g x).re) : by simp only [of_real_one, div_one] - -lemma is_Theta_cpow_rpow (hl_im : is_bounded_under (≤) l (λ x, |(g x).im|)) - (hl : ∀ᶠ x in l, f x = 0 → re (g x) = 0 → g x = 0): - (λ x, f x ^ g x) =Θ[l] (λ x, abs (f x) ^ (g x).re) := -calc (λ x, f x ^ g x) =Θ[l] (λ x, abs (f x) ^ (g x).re / real.exp (arg (f x) * im (g x))) : - is_Theta_of_norm_eventually_eq' $ hl.mono $ λ x, abs_cpow_of_imp -... =Θ[l] (λ x, abs (f x) ^ (g x).re / (1 : ℝ)) : - (is_Theta_refl _ _).div (is_Theta_exp_arg_mul_im hl_im) -... =ᶠ[l] (λ x, abs (f x) ^ (g x).re) : by simp only [of_real_one, div_one] - -lemma is_Theta_cpow_const_rpow {b : ℂ} (hl : b.re = 0 → b ≠ 0 → ∀ᶠ x in l, f x ≠ 0) : - (λ x, f x ^ b) =Θ[l] (λ x, abs (f x) ^ b.re) := -is_Theta_cpow_rpow is_bounded_under_const $ by simpa only [eventually_imp_distrib_right, ne.def, - ← not_frequently, not_imp_not, imp.swap] using hl - -end - -@[simp] lemma abs_cpow_real (x : ℂ) (y : ℝ) : abs (x ^ (y : ℂ)) = x.abs ^ y := -by rcases eq_or_ne x 0 with rfl|hx; [rcases eq_or_ne y 0 with rfl|hy, skip]; - simp [*, abs_cpow_of_ne_zero] - -@[simp] lemma abs_cpow_inv_nat (x : ℂ) (n : ℕ) : abs (x ^ (n⁻¹ : ℂ)) = x.abs ^ (n⁻¹ : ℝ) := -by rw ← abs_cpow_real; simp [-abs_cpow_real] - -lemma abs_cpow_eq_rpow_re_of_pos {x : ℝ} (hx : 0 < x) (y : ℂ) : abs (x ^ y) = x ^ y.re := -by rw [abs_cpow_of_ne_zero (of_real_ne_zero.mpr hx.ne'), arg_of_real_of_nonneg hx.le, zero_mul, - real.exp_zero, div_one, abs_of_nonneg hx.le] - -lemma abs_cpow_eq_rpow_re_of_nonneg {x : ℝ} (hx : 0 ≤ x) {y : ℂ} (hy : re y ≠ 0) : - abs (x ^ y) = x ^ re y := -begin - rcases hx.eq_or_lt with rfl|hlt, - { rw [of_real_zero, zero_cpow, map_zero, real.zero_rpow hy], - exact ne_of_apply_ne re hy }, - { exact abs_cpow_eq_rpow_re_of_pos hlt y } -end - -lemma inv_cpow_eq_ite (x : ℂ) (n : ℂ) : - x⁻¹ ^ n = if x.arg = π then conj (x ^ conj n)⁻¹ else (x ^ n)⁻¹ := -begin - simp_rw [complex.cpow_def, log_inv_eq_ite, inv_eq_zero, map_eq_zero, ite_mul, neg_mul, - is_R_or_C.conj_inv, apply_ite conj, apply_ite exp, apply_ite has_inv.inv, map_zero, map_one, - exp_neg, inv_one, inv_zero, ←exp_conj, map_mul, conj_conj], - split_ifs with hx hn ha ha; refl, -end - -lemma inv_cpow (x : ℂ) (n : ℂ) (hx : x.arg ≠ π) : x⁻¹ ^ n = (x ^ n)⁻¹ := -by rw [inv_cpow_eq_ite, if_neg hx] - -/-- `complex.inv_cpow_eq_ite` with the `ite` on the other side. -/ -lemma inv_cpow_eq_ite' (x : ℂ) (n : ℂ) : - (x ^ n)⁻¹ = if x.arg = π then conj (x⁻¹ ^ conj n) else x⁻¹ ^ n := -begin - rw [inv_cpow_eq_ite, apply_ite conj, conj_conj, conj_conj], - split_ifs, - { refl }, - { rw inv_cpow _ _ h } -end - -lemma conj_cpow_eq_ite (x : ℂ) (n : ℂ) : - conj x ^ n = if x.arg = π then x ^ n else conj (x ^ conj n) := -begin - simp_rw [cpow_def, map_eq_zero, apply_ite conj, map_one, map_zero, ←exp_conj, map_mul, - conj_conj, log_conj_eq_ite], - split_ifs with hcx hn hx; refl -end - -lemma conj_cpow (x : ℂ) (n : ℂ) (hx : x.arg ≠ π) : conj x ^ n = conj (x ^ conj n) := -by rw [conj_cpow_eq_ite, if_neg hx] - -lemma cpow_conj (x : ℂ) (n : ℂ) (hx : x.arg ≠ π) : x ^ conj n = conj (conj x ^ n) := -by rw [conj_cpow _ _ hx, conj_conj] - -end complex - -namespace real - -variables {x y z : ℝ} - -lemma rpow_add (hx : 0 < x) (y z : ℝ) : x ^ (y + z) = x ^ y * x ^ z := -by simp only [rpow_def_of_pos hx, mul_add, exp_add] - -lemma rpow_add' (hx : 0 ≤ x) (h : y + z ≠ 0) : x ^ (y + z) = x ^ y * x ^ z := -begin - rcases hx.eq_or_lt with rfl|pos, - { rw [zero_rpow h, zero_eq_mul], - have : y ≠ 0 ∨ z ≠ 0, from not_and_distrib.1 (λ ⟨hy, hz⟩, h $ hy.symm ▸ hz.symm ▸ zero_add 0), - exact this.imp zero_rpow zero_rpow }, - { exact rpow_add pos _ _ } -end - -lemma rpow_add_of_nonneg (hx : 0 ≤ x) (hy : 0 ≤ y) (hz : 0 ≤ z) : - x ^ (y + z) = x ^ y * x ^ z := -begin - rcases hy.eq_or_lt with rfl|hy, - { rw [zero_add, rpow_zero, one_mul] }, - exact rpow_add' hx (ne_of_gt $ add_pos_of_pos_of_nonneg hy hz) -end - -/-- For `0 ≤ x`, the only problematic case in the equality `x ^ y * x ^ z = x ^ (y + z)` is for -`x = 0` and `y + z = 0`, where the right hand side is `1` while the left hand side can vanish. -The inequality is always true, though, and given in this lemma. -/ -lemma le_rpow_add {x : ℝ} (hx : 0 ≤ x) (y z : ℝ) : x ^ y * x ^ z ≤ x ^ (y + z) := -begin - rcases le_iff_eq_or_lt.1 hx with H|pos, - { by_cases h : y + z = 0, - { simp only [H.symm, h, rpow_zero], - calc (0 : ℝ) ^ y * 0 ^ z ≤ 1 * 1 : - mul_le_mul (zero_rpow_le_one y) (zero_rpow_le_one z) (zero_rpow_nonneg z) zero_le_one - ... = 1 : by simp }, - { simp [rpow_add', ← H, h] } }, - { simp [rpow_add pos] } -end - -lemma rpow_sum_of_pos {ι : Type*} {a : ℝ} (ha : 0 < a) (f : ι → ℝ) (s : finset ι) : - a ^ (∑ x in s, f x) = ∏ x in s, a ^ f x := -@add_monoid_hom.map_sum ℝ ι (additive ℝ) _ _ ⟨λ x : ℝ, (a ^ x : ℝ), rpow_zero a, rpow_add ha⟩ f s - -lemma rpow_sum_of_nonneg {ι : Type*} {a : ℝ} (ha : 0 ≤ a) {s : finset ι} {f : ι → ℝ} - (h : ∀ x ∈ s, 0 ≤ f x) : - a ^ (∑ x in s, f x) = ∏ x in s, a ^ f x := -begin - induction s using finset.cons_induction with i s hi ihs, - { rw [sum_empty, finset.prod_empty, rpow_zero] }, - { rw forall_mem_cons at h, - rw [sum_cons, prod_cons, ← ihs h.2, rpow_add_of_nonneg ha h.1 (sum_nonneg h.2)] } -end - -lemma rpow_mul {x : ℝ} (hx : 0 ≤ x) (y z : ℝ) : x ^ (y * z) = (x ^ y) ^ z := -by rw [← complex.of_real_inj, complex.of_real_cpow (rpow_nonneg_of_nonneg hx _), - complex.of_real_cpow hx, complex.of_real_mul, complex.cpow_mul, complex.of_real_cpow hx]; - simp only [(complex.of_real_mul _ _).symm, (complex.of_real_log hx).symm, - complex.of_real_im, neg_lt_zero, pi_pos, le_of_lt pi_pos] - -lemma rpow_neg {x : ℝ} (hx : 0 ≤ x) (y : ℝ) : x ^ -y = (x ^ y)⁻¹ := -by simp only [rpow_def_of_nonneg hx]; split_ifs; simp [*, exp_neg] at * - -lemma rpow_sub {x : ℝ} (hx : 0 < x) (y z : ℝ) : x ^ (y - z) = x ^ y / x ^ z := -by simp only [sub_eq_add_neg, rpow_add hx, rpow_neg (le_of_lt hx), div_eq_mul_inv] - -lemma rpow_sub' {x : ℝ} (hx : 0 ≤ x) {y z : ℝ} (h : y - z ≠ 0) : - x ^ (y - z) = x ^ y / x ^ z := -by { simp only [sub_eq_add_neg] at h ⊢, simp only [rpow_add' hx h, rpow_neg hx, div_eq_mul_inv] } - -lemma rpow_add_int {x : ℝ} (hx : x ≠ 0) (y : ℝ) (n : ℤ) : x ^ (y + n) = x ^ y * x ^ n := -by rw [rpow_def, complex.of_real_add, complex.cpow_add _ _ (complex.of_real_ne_zero.mpr hx), - complex.of_real_int_cast, complex.cpow_int_cast, ← complex.of_real_zpow, mul_comm, - complex.of_real_mul_re, ← rpow_def, mul_comm] - -lemma rpow_add_nat {x : ℝ} (hx : x ≠ 0) (y : ℝ) (n : ℕ) : x ^ (y + n) = x ^ y * x ^ n := -by simpa using rpow_add_int hx y n - -lemma rpow_sub_int {x : ℝ} (hx : x ≠ 0) (y : ℝ) (n : ℤ) : x ^ (y - n) = x ^ y / x ^ n := -by simpa using rpow_add_int hx y (-n) - -lemma rpow_sub_nat {x : ℝ} (hx : x ≠ 0) (y : ℝ) (n : ℕ) : x ^ (y - n) = x ^ y / x ^ n := -by simpa using rpow_sub_int hx y n - -lemma rpow_add_one {x : ℝ} (hx : x ≠ 0) (y : ℝ) : x ^ (y + 1) = x ^ y * x := -by simpa using rpow_add_nat hx y 1 - -lemma rpow_sub_one {x : ℝ} (hx : x ≠ 0) (y : ℝ) : x ^ (y - 1) = x ^ y / x := -by simpa using rpow_sub_nat hx y 1 - -@[simp, norm_cast] lemma rpow_int_cast (x : ℝ) (n : ℤ) : x ^ (n : ℝ) = x ^ n := -by simp only [rpow_def, ← complex.of_real_zpow, complex.cpow_int_cast, - complex.of_real_int_cast, complex.of_real_re] - -@[simp, norm_cast] lemma rpow_nat_cast (x : ℝ) (n : ℕ) : x ^ (n : ℝ) = x ^ n := -by simpa using rpow_int_cast x n - -@[simp] lemma rpow_two (x : ℝ) : x ^ (2 : ℝ) = x ^ 2 := -by { rw ← rpow_nat_cast, simp only [nat.cast_bit0, nat.cast_one] } - -lemma rpow_neg_one (x : ℝ) : x ^ (-1 : ℝ) = x⁻¹ := -begin - suffices H : x ^ ((-1 : ℤ) : ℝ) = x⁻¹, by rwa [int.cast_neg, int.cast_one] at H, - simp only [rpow_int_cast, zpow_one, zpow_neg], -end - -lemma mul_rpow {x y z : ℝ} (h : 0 ≤ x) (h₁ : 0 ≤ y) : (x*y)^z = x^z * y^z := -begin - iterate 3 { rw real.rpow_def_of_nonneg }, split_ifs; simp * at *, - { have hx : 0 < x, - { cases lt_or_eq_of_le h with h₂ h₂, { exact h₂ }, - exfalso, apply h_2, exact eq.symm h₂ }, - have hy : 0 < y, - { cases lt_or_eq_of_le h₁ with h₂ h₂, { exact h₂ }, - exfalso, apply h_3, exact eq.symm h₂ }, - rw [log_mul (ne_of_gt hx) (ne_of_gt hy), add_mul, exp_add]}, - { exact h₁ }, - { exact h }, - { exact mul_nonneg h h₁ }, -end - -lemma inv_rpow (hx : 0 ≤ x) (y : ℝ) : (x⁻¹)^y = (x^y)⁻¹ := -by simp only [← rpow_neg_one, ← rpow_mul hx, mul_comm] - -lemma div_rpow (hx : 0 ≤ x) (hy : 0 ≤ y) (z : ℝ) : (x / y) ^ z = x^z / y^z := -by simp only [div_eq_mul_inv, mul_rpow hx (inv_nonneg.2 hy), inv_rpow hy] - -lemma log_rpow {x : ℝ} (hx : 0 < x) (y : ℝ) : log (x^y) = y * (log x) := -begin - apply exp_injective, - rw [exp_log (rpow_pos_of_pos hx y), ← exp_log hx, mul_comm, rpow_def_of_pos (exp_pos (log x)) y], -end - -lemma rpow_lt_rpow (hx : 0 ≤ x) (hxy : x < y) (hz : 0 < z) : x^z < y^z := -begin - rw le_iff_eq_or_lt at hx, cases hx, - { rw [← hx, zero_rpow (ne_of_gt hz)], exact rpow_pos_of_pos (by rwa ← hx at hxy) _ }, - rw [rpow_def_of_pos hx, rpow_def_of_pos (lt_trans hx hxy), exp_lt_exp], - exact mul_lt_mul_of_pos_right (log_lt_log hx hxy) hz -end - -lemma rpow_le_rpow {x y z: ℝ} (h : 0 ≤ x) (h₁ : x ≤ y) (h₂ : 0 ≤ z) : x^z ≤ y^z := -begin - rcases eq_or_lt_of_le h₁ with rfl|h₁', { refl }, - rcases eq_or_lt_of_le h₂ with rfl|h₂', { simp }, - exact le_of_lt (rpow_lt_rpow h h₁' h₂') -end - -lemma rpow_lt_rpow_iff (hx : 0 ≤ x) (hy : 0 ≤ y) (hz : 0 < z) : x ^ z < y ^ z ↔ x < y := -⟨lt_imp_lt_of_le_imp_le $ λ h, rpow_le_rpow hy h (le_of_lt hz), λ h, rpow_lt_rpow hx h hz⟩ - -lemma rpow_le_rpow_iff (hx : 0 ≤ x) (hy : 0 ≤ y) (hz : 0 < z) : x ^ z ≤ y ^ z ↔ x ≤ y := -le_iff_le_iff_lt_iff_lt.2 $ rpow_lt_rpow_iff hy hx hz - -lemma le_rpow_inv_iff_of_neg (hx : 0 < x) (hy : 0 < y) (hz : z < 0) : - x ≤ y ^ z⁻¹ ↔ y ≤ x ^ z := -begin - have hz' : 0 < -z := by rwa [lt_neg, neg_zero], - have hxz : 0 < x ^ (-z) := real.rpow_pos_of_pos hx _, - have hyz : 0 < y ^ z⁻¹ := real.rpow_pos_of_pos hy _, - rw [←real.rpow_le_rpow_iff hx.le hyz.le hz', ←real.rpow_mul hy.le], - simp only [ne_of_lt hz, real.rpow_neg_one, mul_neg, inv_mul_cancel, ne.def, not_false_iff], - rw [le_inv hxz hy, ←real.rpow_neg_one, ←real.rpow_mul hx.le], - simp, -end - -lemma lt_rpow_inv_iff_of_neg (hx : 0 < x) (hy : 0 < y) (hz : z < 0) : - x < y ^ z⁻¹ ↔ y < x ^ z := -begin - have hz' : 0 < -z := by rwa [lt_neg, neg_zero], - have hxz : 0 < x ^ (-z) := real.rpow_pos_of_pos hx _, - have hyz : 0 < y ^ z⁻¹ := real.rpow_pos_of_pos hy _, - rw [←real.rpow_lt_rpow_iff hx.le hyz.le hz', ←real.rpow_mul hy.le], - simp only [ne_of_lt hz, real.rpow_neg_one, mul_neg, inv_mul_cancel, ne.def, not_false_iff], - rw [lt_inv hxz hy, ←real.rpow_neg_one, ←real.rpow_mul hx.le], - simp, -end - -lemma rpow_inv_lt_iff_of_neg (hx : 0 < x) (hy : 0 < y) (hz : z < 0) : - x ^ z⁻¹ < y ↔ y ^ z < x := -begin - convert lt_rpow_inv_iff_of_neg (real.rpow_pos_of_pos hx _) (real.rpow_pos_of_pos hy _) hz; - simp [←real.rpow_mul hx.le, ←real.rpow_mul hy.le, ne_of_lt hz], -end - -lemma rpow_inv_le_iff_of_neg (hx : 0 < x) (hy : 0 < y) (hz : z < 0) : - x ^ z⁻¹ ≤ y ↔ y ^ z ≤ x := -begin - convert le_rpow_inv_iff_of_neg (real.rpow_pos_of_pos hx _) (real.rpow_pos_of_pos hy _) hz; - simp [←real.rpow_mul hx.le, ←real.rpow_mul hy.le, ne_of_lt hz], -end - -lemma rpow_lt_rpow_of_exponent_lt (hx : 1 < x) (hyz : y < z) : x^y < x^z := -begin - repeat {rw [rpow_def_of_pos (lt_trans zero_lt_one hx)]}, - rw exp_lt_exp, exact mul_lt_mul_of_pos_left hyz (log_pos hx), -end - -lemma rpow_le_rpow_of_exponent_le (hx : 1 ≤ x) (hyz : y ≤ z) : x^y ≤ x^z := -begin - repeat {rw [rpow_def_of_pos (lt_of_lt_of_le zero_lt_one hx)]}, - rw exp_le_exp, exact mul_le_mul_of_nonneg_left hyz (log_nonneg hx), -end - -@[simp] lemma rpow_le_rpow_left_iff (hx : 1 < x) : x ^ y ≤ x ^ z ↔ y ≤ z := -begin - have x_pos : 0 < x := lt_trans zero_lt_one hx, - rw [←log_le_log (rpow_pos_of_pos x_pos y) (rpow_pos_of_pos x_pos z), - log_rpow x_pos, log_rpow x_pos, mul_le_mul_right (log_pos hx)], -end - -@[simp] lemma rpow_lt_rpow_left_iff (hx : 1 < x) : x ^ y < x ^ z ↔ y < z := -by rw [lt_iff_not_le, rpow_le_rpow_left_iff hx, lt_iff_not_le] - -lemma rpow_lt_rpow_of_exponent_gt (hx0 : 0 < x) (hx1 : x < 1) (hyz : z < y) : - x^y < x^z := -begin - repeat {rw [rpow_def_of_pos hx0]}, - rw exp_lt_exp, exact mul_lt_mul_of_neg_left hyz (log_neg hx0 hx1), -end - -lemma rpow_le_rpow_of_exponent_ge (hx0 : 0 < x) (hx1 : x ≤ 1) (hyz : z ≤ y) : - x^y ≤ x^z := -begin - repeat {rw [rpow_def_of_pos hx0]}, - rw exp_le_exp, exact mul_le_mul_of_nonpos_left hyz (log_nonpos (le_of_lt hx0) hx1), -end - -@[simp] lemma rpow_le_rpow_left_iff_of_base_lt_one (hx0 : 0 < x) (hx1 : x < 1) : - x ^ y ≤ x ^ z ↔ z ≤ y := -begin - rw [←log_le_log (rpow_pos_of_pos hx0 y) (rpow_pos_of_pos hx0 z), - log_rpow hx0, log_rpow hx0, mul_le_mul_right_of_neg (log_neg hx0 hx1)], -end - -@[simp] lemma rpow_lt_rpow_left_iff_of_base_lt_one (hx0 : 0 < x) (hx1 : x < 1) : - x ^ y < x ^ z ↔ z < y := -by rw [lt_iff_not_le, rpow_le_rpow_left_iff_of_base_lt_one hx0 hx1, lt_iff_not_le] - -lemma rpow_lt_one {x z : ℝ} (hx1 : 0 ≤ x) (hx2 : x < 1) (hz : 0 < z) : x^z < 1 := -by { rw ← one_rpow z, exact rpow_lt_rpow hx1 hx2 hz } - -lemma rpow_le_one {x z : ℝ} (hx1 : 0 ≤ x) (hx2 : x ≤ 1) (hz : 0 ≤ z) : x^z ≤ 1 := -by { rw ← one_rpow z, exact rpow_le_rpow hx1 hx2 hz } - -lemma rpow_lt_one_of_one_lt_of_neg {x z : ℝ} (hx : 1 < x) (hz : z < 0) : x^z < 1 := -by { convert rpow_lt_rpow_of_exponent_lt hx hz, exact (rpow_zero x).symm } - -lemma rpow_le_one_of_one_le_of_nonpos {x z : ℝ} (hx : 1 ≤ x) (hz : z ≤ 0) : x^z ≤ 1 := -by { convert rpow_le_rpow_of_exponent_le hx hz, exact (rpow_zero x).symm } - -lemma one_lt_rpow {x z : ℝ} (hx : 1 < x) (hz : 0 < z) : 1 < x^z := -by { rw ← one_rpow z, exact rpow_lt_rpow zero_le_one hx hz } - -lemma one_le_rpow {x z : ℝ} (hx : 1 ≤ x) (hz : 0 ≤ z) : 1 ≤ x^z := -by { rw ← one_rpow z, exact rpow_le_rpow zero_le_one hx hz } - -lemma one_lt_rpow_of_pos_of_lt_one_of_neg (hx1 : 0 < x) (hx2 : x < 1) (hz : z < 0) : - 1 < x^z := -by { convert rpow_lt_rpow_of_exponent_gt hx1 hx2 hz, exact (rpow_zero x).symm } - -lemma one_le_rpow_of_pos_of_le_one_of_nonpos (hx1 : 0 < x) (hx2 : x ≤ 1) (hz : z ≤ 0) : - 1 ≤ x^z := -by { convert rpow_le_rpow_of_exponent_ge hx1 hx2 hz, exact (rpow_zero x).symm } - -lemma rpow_lt_one_iff_of_pos (hx : 0 < x) : x ^ y < 1 ↔ 1 < x ∧ y < 0 ∨ x < 1 ∧ 0 < y := -by rw [rpow_def_of_pos hx, exp_lt_one_iff, mul_neg_iff, log_pos_iff hx, log_neg_iff hx] - -lemma rpow_lt_one_iff (hx : 0 ≤ x) : x ^ y < 1 ↔ x = 0 ∧ y ≠ 0 ∨ 1 < x ∧ y < 0 ∨ x < 1 ∧ 0 < y := -begin - rcases hx.eq_or_lt with (rfl|hx), - { rcases em (y = 0) with (rfl|hy); simp [*, lt_irrefl, zero_lt_one] }, - { simp [rpow_lt_one_iff_of_pos hx, hx.ne.symm] } -end - -lemma one_lt_rpow_iff_of_pos (hx : 0 < x) : 1 < x ^ y ↔ 1 < x ∧ 0 < y ∨ x < 1 ∧ y < 0 := -by rw [rpow_def_of_pos hx, one_lt_exp_iff, mul_pos_iff, log_pos_iff hx, log_neg_iff hx] - -lemma one_lt_rpow_iff (hx : 0 ≤ x) : 1 < x ^ y ↔ 1 < x ∧ 0 < y ∨ 0 < x ∧ x < 1 ∧ y < 0 := -begin - rcases hx.eq_or_lt with (rfl|hx), - { rcases em (y = 0) with (rfl|hy); simp [*, lt_irrefl, (zero_lt_one' ℝ).not_lt] }, - { simp [one_lt_rpow_iff_of_pos hx, hx] } -end - -lemma rpow_le_rpow_of_exponent_ge' (hx0 : 0 ≤ x) (hx1 : x ≤ 1) (hz : 0 ≤ z) (hyz : z ≤ y) : - x^y ≤ x^z := -begin - rcases eq_or_lt_of_le hx0 with rfl | hx0', - { rcases eq_or_lt_of_le hz with rfl | hz', - { exact (rpow_zero 0).symm ▸ (rpow_le_one hx0 hx1 hyz), }, - rw [zero_rpow, zero_rpow]; linarith, }, - { exact rpow_le_rpow_of_exponent_ge hx0' hx1 hyz, }, -end - -lemma rpow_left_inj_on {x : ℝ} (hx : x ≠ 0) : - inj_on (λ y : ℝ, y^x) {y : ℝ | 0 ≤ y} := -begin - rintros y hy z hz (hyz : y ^ x = z ^ x), - rw [←rpow_one y, ←rpow_one z, ←_root_.mul_inv_cancel hx, rpow_mul hy, rpow_mul hz, hyz] -end - -lemma le_rpow_iff_log_le (hx : 0 < x) (hy : 0 < y) : - x ≤ y^z ↔ real.log x ≤ z * real.log y := -by rw [←real.log_le_log hx (real.rpow_pos_of_pos hy z), real.log_rpow hy] - -lemma le_rpow_of_log_le (hx : 0 ≤ x) (hy : 0 < y) (h : real.log x ≤ z * real.log y) : - x ≤ y^z := -begin - obtain hx | rfl := hx.lt_or_eq, - { exact (le_rpow_iff_log_le hx hy).2 h }, - exact (real.rpow_pos_of_pos hy z).le, -end - -lemma lt_rpow_iff_log_lt (hx : 0 < x) (hy : 0 < y) : - x < y^z ↔ real.log x < z * real.log y := -by rw [←real.log_lt_log_iff hx (real.rpow_pos_of_pos hy z), real.log_rpow hy] - -lemma lt_rpow_of_log_lt (hx : 0 ≤ x) (hy : 0 < y) (h : real.log x < z * real.log y) : - x < y^z := -begin - obtain hx | rfl := hx.lt_or_eq, - { exact (lt_rpow_iff_log_lt hx hy).2 h }, - exact real.rpow_pos_of_pos hy z, -end - -lemma rpow_le_one_iff_of_pos (hx : 0 < x) : x ^ y ≤ 1 ↔ 1 ≤ x ∧ y ≤ 0 ∨ x ≤ 1 ∧ 0 ≤ y := -by rw [rpow_def_of_pos hx, exp_le_one_iff, mul_nonpos_iff, log_nonneg_iff hx, log_nonpos_iff hx] - -/-- Bound for `|log x * x ^ t|` in the interval `(0, 1]`, for positive real `t`. -/ -lemma abs_log_mul_self_rpow_lt (x t : ℝ) (h1 : 0 < x) (h2 : x ≤ 1) (ht : 0 < t) : - |log x * x ^ t| < 1 / t := -begin - rw lt_div_iff ht, - have := abs_log_mul_self_lt (x ^ t) (rpow_pos_of_pos h1 t) (rpow_le_one h1.le h2 ht.le), - rwa [log_rpow h1, mul_assoc, abs_mul, abs_of_pos ht, mul_comm] at this -end - -lemma pow_nat_rpow_nat_inv {x : ℝ} (hx : 0 ≤ x) {n : ℕ} (hn : n ≠ 0) : - (x ^ n) ^ (n⁻¹ : ℝ) = x := -have hn0 : (n : ℝ) ≠ 0, from nat.cast_ne_zero.2 hn, -by rw [← rpow_nat_cast, ← rpow_mul hx, mul_inv_cancel hn0, rpow_one] - -lemma rpow_nat_inv_pow_nat {x : ℝ} (hx : 0 ≤ x) {n : ℕ} (hn : n ≠ 0) : - (x ^ (n⁻¹ : ℝ)) ^ n = x := -have hn0 : (n : ℝ) ≠ 0, from nat.cast_ne_zero.2 hn, -by rw [← rpow_nat_cast, ← rpow_mul hx, inv_mul_cancel hn0, rpow_one] - -lemma continuous_at_const_rpow {a b : ℝ} (h : a ≠ 0) : continuous_at (rpow a) b := -begin - have : rpow a = λ x : ℝ, ((a : ℂ) ^ (x : ℂ)).re, by { ext1 x, rw [rpow_eq_pow, rpow_def], }, - rw this, - refine complex.continuous_re.continuous_at.comp _, - refine (continuous_at_const_cpow _).comp complex.continuous_of_real.continuous_at, - norm_cast, - exact h, -end - -lemma continuous_at_const_rpow' {a b : ℝ} (h : b ≠ 0) : continuous_at (rpow a) b := -begin - have : rpow a = λ x : ℝ, ((a : ℂ) ^ (x : ℂ)).re, by { ext1 x, rw [rpow_eq_pow, rpow_def], }, - rw this, - refine complex.continuous_re.continuous_at.comp _, - refine (continuous_at_const_cpow' _).comp complex.continuous_of_real.continuous_at, - norm_cast, - exact h, -end - -lemma rpow_eq_nhds_of_neg {p : ℝ × ℝ} (hp_fst : p.fst < 0) : - (λ x : ℝ × ℝ, x.1 ^ x.2) =ᶠ[𝓝 p] λ x, exp (log x.1 * x.2) * cos (x.2 * π) := -begin - suffices : ∀ᶠ (x : ℝ × ℝ) in (𝓝 p), x.1 < 0, - from this.mono (λ x hx, by { dsimp only, rw rpow_def_of_neg hx, }), - exact is_open.eventually_mem (is_open_lt continuous_fst continuous_const) hp_fst, -end - -lemma rpow_eq_nhds_of_pos {p : ℝ × ℝ} (hp_fst : 0 < p.fst) : - (λ x : ℝ × ℝ, x.1 ^ x.2) =ᶠ[𝓝 p] λ x, exp (log x.1 * x.2) := -begin - suffices : ∀ᶠ (x : ℝ × ℝ) in (𝓝 p), 0 < x.1, - from this.mono (λ x hx, by { dsimp only, rw rpow_def_of_pos hx, }), - exact is_open.eventually_mem (is_open_lt continuous_const continuous_fst) hp_fst, -end - -lemma continuous_at_rpow_of_ne (p : ℝ × ℝ) (hp : p.1 ≠ 0) : - continuous_at (λ p : ℝ × ℝ, p.1 ^ p.2) p := -begin - rw ne_iff_lt_or_gt at hp, - cases hp, - { rw continuous_at_congr (rpow_eq_nhds_of_neg hp), - refine continuous_at.mul _ (continuous_cos.continuous_at.comp _), - { refine continuous_exp.continuous_at.comp (continuous_at.mul _ continuous_snd.continuous_at), - refine (continuous_at_log _).comp continuous_fst.continuous_at, - exact hp.ne, }, - { exact continuous_snd.continuous_at.mul continuous_at_const, }, }, - { rw continuous_at_congr (rpow_eq_nhds_of_pos hp), - refine continuous_exp.continuous_at.comp (continuous_at.mul _ continuous_snd.continuous_at), - refine (continuous_at_log _).comp continuous_fst.continuous_at, - exact hp.lt.ne.symm, }, -end - -lemma continuous_at_rpow_of_pos (p : ℝ × ℝ) (hp : 0 < p.2) : - continuous_at (λ p : ℝ × ℝ, p.1 ^ p.2) p := -begin - cases p with x y, - obtain hx|rfl := ne_or_eq x 0, - { exact continuous_at_rpow_of_ne (x, y) hx }, - have A : tendsto (λ p : ℝ × ℝ, exp (log p.1 * p.2)) (𝓝[≠] 0 ×ᶠ 𝓝 y) (𝓝 0) := - tendsto_exp_at_bot.comp - ((tendsto_log_nhds_within_zero.comp tendsto_fst).at_bot_mul hp tendsto_snd), - have B : tendsto (λ p : ℝ × ℝ, p.1 ^ p.2) (𝓝[≠] 0 ×ᶠ 𝓝 y) (𝓝 0) := - squeeze_zero_norm (λ p, abs_rpow_le_exp_log_mul p.1 p.2) A, - have C : tendsto (λ p : ℝ × ℝ, p.1 ^ p.2) (𝓝[{0}] 0 ×ᶠ 𝓝 y) (pure 0), - { rw [nhds_within_singleton, tendsto_pure, pure_prod, eventually_map], - exact (lt_mem_nhds hp).mono (λ y hy, zero_rpow hy.ne') }, - simpa only [← sup_prod, ← nhds_within_union, compl_union_self, nhds_within_univ, nhds_prod_eq, - continuous_at, zero_rpow hp.ne'] using B.sup (C.mono_right (pure_le_nhds _)) -end - -lemma continuous_at_rpow (p : ℝ × ℝ) (h : p.1 ≠ 0 ∨ 0 < p.2) : - continuous_at (λ p : ℝ × ℝ, p.1 ^ p.2) p := -h.elim (λ h, continuous_at_rpow_of_ne p h) (λ h, continuous_at_rpow_of_pos p h) - -lemma continuous_at_rpow_const (x : ℝ) (q : ℝ) (h : x ≠ 0 ∨ 0 < q) : - continuous_at (λ (x : ℝ), x ^ q) x := -begin - change continuous_at ((λ p : ℝ × ℝ, p.1 ^ p.2) ∘ (λ y : ℝ, (y, q))) x, - apply continuous_at.comp, - { exact continuous_at_rpow (x, q) h }, - { exact (continuous_id'.prod_mk continuous_const).continuous_at } -end - -end real - -section - -variable {α : Type*} - -lemma filter.tendsto.rpow {l : filter α} {f g : α → ℝ} {x y : ℝ} - (hf : tendsto f l (𝓝 x)) (hg : tendsto g l (𝓝 y)) (h : x ≠ 0 ∨ 0 < y) : - tendsto (λ t, f t ^ g t) l (𝓝 (x ^ y)) := -(real.continuous_at_rpow (x, y) h).tendsto.comp (hf.prod_mk_nhds hg) - -lemma filter.tendsto.rpow_const {l : filter α} {f : α → ℝ} {x p : ℝ} - (hf : tendsto f l (𝓝 x)) (h : x ≠ 0 ∨ 0 ≤ p) : - tendsto (λ a, f a ^ p) l (𝓝 (x ^ p)) := -if h0 : 0 = p then h0 ▸ by simp [tendsto_const_nhds] -else hf.rpow tendsto_const_nhds (h.imp id $ λ h', h'.lt_of_ne h0) - -variables [topological_space α] {f g : α → ℝ} {s : set α} {x : α} {p : ℝ} - -lemma continuous_at.rpow (hf : continuous_at f x) (hg : continuous_at g x) (h : f x ≠ 0 ∨ 0 < g x) : - continuous_at (λ t, f t ^ g t) x := -hf.rpow hg h - -lemma continuous_within_at.rpow (hf : continuous_within_at f s x) (hg : continuous_within_at g s x) - (h : f x ≠ 0 ∨ 0 < g x) : - continuous_within_at (λ t, f t ^ g t) s x := -hf.rpow hg h - -lemma continuous_on.rpow (hf : continuous_on f s) (hg : continuous_on g s) - (h : ∀ x ∈ s, f x ≠ 0 ∨ 0 < g x) : - continuous_on (λ t, f t ^ g t) s := -λ t ht, (hf t ht).rpow (hg t ht) (h t ht) - -lemma continuous.rpow (hf : continuous f) (hg : continuous g) (h : ∀ x, f x ≠ 0 ∨ 0 < g x) : - continuous (λ x, f x ^ g x) := -continuous_iff_continuous_at.2 $ λ x, (hf.continuous_at.rpow hg.continuous_at (h x)) - -lemma continuous_within_at.rpow_const (hf : continuous_within_at f s x) (h : f x ≠ 0 ∨ 0 ≤ p) : - continuous_within_at (λ x, f x ^ p) s x := -hf.rpow_const h - -lemma continuous_at.rpow_const (hf : continuous_at f x) (h : f x ≠ 0 ∨ 0 ≤ p) : - continuous_at (λ x, f x ^ p) x := -hf.rpow_const h - -lemma continuous_on.rpow_const (hf : continuous_on f s) (h : ∀ x ∈ s, f x ≠ 0 ∨ 0 ≤ p) : - continuous_on (λ x, f x ^ p) s := -λ x hx, (hf x hx).rpow_const (h x hx) - -lemma continuous.rpow_const (hf : continuous f) (h : ∀ x, f x ≠ 0 ∨ 0 ≤ p) : - continuous (λ x, f x ^ p) := -continuous_iff_continuous_at.2 $ λ x, hf.continuous_at.rpow_const (h x) - -end - -namespace real - -variables {z x y : ℝ} - -section sqrt - -lemma sqrt_eq_rpow (x : ℝ) : sqrt x = x ^ (1/(2:ℝ)) := -begin - obtain h | h := le_or_lt 0 x, - { rw [← mul_self_inj_of_nonneg (sqrt_nonneg _) (rpow_nonneg_of_nonneg h _), mul_self_sqrt h, - ← sq, ← rpow_nat_cast, ← rpow_mul h], - norm_num }, - { have : 1 / (2:ℝ) * π = π / (2:ℝ), ring, - rw [sqrt_eq_zero_of_nonpos h.le, rpow_def_of_neg h, this, cos_pi_div_two, mul_zero] } -end - -lemma rpow_div_two_eq_sqrt {x : ℝ} (r : ℝ) (hx : 0 ≤ x) : x ^ (r/2) = (sqrt x) ^ r := -begin - rw [sqrt_eq_rpow, ← rpow_mul hx], - congr, - ring, -end - -end sqrt - -end real - -section limits -open real filter - -/-- The function `x ^ y` tends to `+∞` at `+∞` for any positive real `y`. -/ -lemma tendsto_rpow_at_top {y : ℝ} (hy : 0 < y) : tendsto (λ x : ℝ, x ^ y) at_top at_top := -begin - rw tendsto_at_top_at_top, - intro b, - use (max b 0) ^ (1/y), - intros x hx, - exact le_of_max_le_left - (by { convert rpow_le_rpow (rpow_nonneg_of_nonneg (le_max_right b 0) (1/y)) hx (le_of_lt hy), - rw [← rpow_mul (le_max_right b 0), (eq_div_iff (ne_of_gt hy)).mp rfl, rpow_one] }), -end - -/-- The function `x ^ (-y)` tends to `0` at `+∞` for any positive real `y`. -/ -lemma tendsto_rpow_neg_at_top {y : ℝ} (hy : 0 < y) : tendsto (λ x : ℝ, x ^ (-y)) at_top (𝓝 0) := -tendsto.congr' (eventually_eq_of_mem (Ioi_mem_at_top 0) (λ x hx, (rpow_neg (le_of_lt hx) y).symm)) - (tendsto_rpow_at_top hy).inv_tendsto_at_top - -/-- The function `x ^ (a / (b * x + c))` tends to `1` at `+∞`, for any real numbers `a`, `b`, and -`c` such that `b` is nonzero. -/ -lemma tendsto_rpow_div_mul_add (a b c : ℝ) (hb : 0 ≠ b) : - tendsto (λ x, x ^ (a / (b*x+c))) at_top (𝓝 1) := -begin - refine tendsto.congr' _ ((tendsto_exp_nhds_0_nhds_1.comp - (by simpa only [mul_zero, pow_one] using ((@tendsto_const_nhds _ _ _ a _).mul - (tendsto_div_pow_mul_exp_add_at_top b c 1 hb)))).comp tendsto_log_at_top), - apply eventually_eq_of_mem (Ioi_mem_at_top (0:ℝ)), - intros x hx, - simp only [set.mem_Ioi, function.comp_app] at hx ⊢, - rw [exp_log hx, ← exp_log (rpow_pos_of_pos hx (a / (b * x + c))), log_rpow hx (a / (b * x + c))], - field_simp, -end - -/-- The function `x ^ (1 / x)` tends to `1` at `+∞`. -/ -lemma tendsto_rpow_div : tendsto (λ x, x ^ ((1:ℝ) / x)) at_top (𝓝 1) := -by { convert tendsto_rpow_div_mul_add (1:ℝ) _ (0:ℝ) zero_ne_one, funext, congr' 2, ring } - -/-- The function `x ^ (-1 / x)` tends to `1` at `+∞`. -/ -lemma tendsto_rpow_neg_div : tendsto (λ x, x ^ (-(1:ℝ) / x)) at_top (𝓝 1) := -by { convert tendsto_rpow_div_mul_add (-(1:ℝ)) _ (0:ℝ) zero_ne_one, funext, congr' 2, ring } - -/-- The function `exp(x) / x ^ s` tends to `+∞` at `+∞`, for any real number `s`. -/ -lemma tendsto_exp_div_rpow_at_top (s : ℝ) : tendsto (λ x : ℝ, exp x / x ^ s) at_top at_top := -begin - cases archimedean_iff_nat_lt.1 (real.archimedean) s with n hn, - refine tendsto_at_top_mono' _ _ (tendsto_exp_div_pow_at_top n), - filter_upwards [eventually_gt_at_top (0 : ℝ), eventually_ge_at_top (1 : ℝ)] with x hx₀ hx₁, - rw [div_le_div_left (exp_pos _) (pow_pos hx₀ _) (rpow_pos_of_pos hx₀ _), ←rpow_nat_cast], - exact rpow_le_rpow_of_exponent_le hx₁ hn.le, -end - -/-- The function `exp (b * x) / x ^ s` tends to `+∞` at `+∞`, for any real `s` and `b > 0`. -/ -lemma tendsto_exp_mul_div_rpow_at_top (s : ℝ) (b : ℝ) (hb : 0 < b) : - tendsto (λ x : ℝ, exp (b * x) / x ^ s) at_top at_top := -begin - refine ((tendsto_rpow_at_top hb).comp (tendsto_exp_div_rpow_at_top (s / b))).congr' _, - filter_upwards [eventually_ge_at_top (0 : ℝ)] with x hx₀, - simp [div_rpow, (exp_pos x).le, rpow_nonneg_of_nonneg, ←rpow_mul, ←exp_mul, mul_comm x, hb.ne', *] -end - -/-- The function `x ^ s * exp (-b * x)` tends to `0` at `+∞`, for any real `s` and `b > 0`. -/ -lemma tendsto_rpow_mul_exp_neg_mul_at_top_nhds_0 (s : ℝ) (b : ℝ) (hb : 0 < b): - tendsto (λ x : ℝ, x ^ s * exp (-b * x)) at_top (𝓝 0) := -begin - refine (tendsto_exp_mul_div_rpow_at_top s b hb).inv_tendsto_at_top.congr' _, - filter_upwards with x using by simp [exp_neg, inv_div, div_eq_mul_inv _ (exp _)] -end - -namespace asymptotics - -variables {α : Type*} {r c : ℝ} {l : filter α} {f g : α → ℝ} - -lemma is_O_with.rpow (h : is_O_with c l f g) (hc : 0 ≤ c) (hr : 0 ≤ r) (hg : 0 ≤ᶠ[l] g) : - is_O_with (c ^ r) l (λ x, f x ^ r) (λ x, g x ^ r) := -begin - apply is_O_with.of_bound, - filter_upwards [hg, h.bound] with x hgx hx, - calc |f x ^ r| ≤ |f x| ^ r : abs_rpow_le_abs_rpow _ _ - ... ≤ (c * |g x|) ^ r : rpow_le_rpow (abs_nonneg _) hx hr - ... = c ^ r * |g x ^ r| : by rw [mul_rpow hc (abs_nonneg _), abs_rpow_of_nonneg hgx] -end - -lemma is_O.rpow (hr : 0 ≤ r) (hg : 0 ≤ᶠ[l] g) (h : f =O[l] g) : - (λ x, f x ^ r) =O[l] (λ x, g x ^ r) := -let ⟨c, hc, h'⟩ := h.exists_nonneg in (h'.rpow hc hr hg).is_O - -lemma is_o.rpow (hr : 0 < r) (hg : 0 ≤ᶠ[l] g) (h : f =o[l] g) : - (λ x, f x ^ r) =o[l] (λ x, g x ^ r) := -is_o.of_is_O_with $ λ c hc, ((h.forall_is_O_with (rpow_pos_of_pos hc r⁻¹)).rpow - (rpow_nonneg_of_nonneg hc.le _) hr.le hg).congr_const - (by rw [←rpow_mul hc.le, inv_mul_cancel hr.ne', rpow_one]) - -end asymptotics - -open asymptotics - -/-- `x ^ s = o(exp(b * x))` as `x → ∞` for any real `s` and positive `b`. -/ -lemma is_o_rpow_exp_pos_mul_at_top (s : ℝ) {b : ℝ} (hb : 0 < b) : - (λ x : ℝ, x ^ s) =o[at_top] (λ x, exp (b * x)) := -iff.mpr (is_o_iff_tendsto $ λ x h, absurd h (exp_pos _).ne') $ - by simpa only [div_eq_mul_inv, exp_neg, neg_mul] - using tendsto_rpow_mul_exp_neg_mul_at_top_nhds_0 s b hb - -/-- `x ^ k = o(exp(b * x))` as `x → ∞` for any integer `k` and positive `b`. -/ -lemma is_o_zpow_exp_pos_mul_at_top (k : ℤ) {b : ℝ} (hb : 0 < b) : - (λ x : ℝ, x ^ k) =o[at_top] (λ x, exp (b * x)) := -by simpa only [rpow_int_cast] using is_o_rpow_exp_pos_mul_at_top k hb - -/-- `x ^ k = o(exp(b * x))` as `x → ∞` for any natural `k` and positive `b`. -/ -lemma is_o_pow_exp_pos_mul_at_top (k : ℕ) {b : ℝ} (hb : 0 < b) : - (λ x : ℝ, x ^ k) =o[at_top] (λ x, exp (b * x)) := -by simpa using is_o_zpow_exp_pos_mul_at_top k hb - -/-- `x ^ s = o(exp x)` as `x → ∞` for any real `s`. -/ -lemma is_o_rpow_exp_at_top (s : ℝ) : (λ x : ℝ, x ^ s) =o[at_top] exp := -by simpa only [one_mul] using is_o_rpow_exp_pos_mul_at_top s one_pos - -/-- `exp (-a * x) = o(x ^ s)` as `x → ∞`, for any positive `a` and real `s`. -/ -lemma is_o_exp_neg_mul_rpow_at_top {a : ℝ} (ha : 0 < a) (b : ℝ) : - is_o at_top (λ x : ℝ, exp (-a * x)) (λ x : ℝ, x ^ b) := -begin - apply is_o_of_tendsto', - { refine (eventually_gt_at_top 0).mp (eventually_of_forall $ λ t ht h, _), - rw rpow_eq_zero_iff_of_nonneg ht.le at h, - exact (ht.ne' h.1).elim }, - { refine (tendsto_exp_mul_div_rpow_at_top (-b) a ha).inv_tendsto_at_top.congr' _, - refine (eventually_ge_at_top 0).mp (eventually_of_forall $ λ t ht, _), - dsimp only, - rw [pi.inv_apply, inv_div, ←inv_div_inv, neg_mul, real.exp_neg, rpow_neg ht, inv_inv] } -end - -lemma is_o_log_rpow_at_top {r : ℝ} (hr : 0 < r) : log =o[at_top] (λ x, x ^ r) := -calc log =O[at_top] (λ x, r * log x) : is_O_self_const_mul _ hr.ne' _ _ - ... =ᶠ[at_top] (λ x, log (x ^ r)) : - (eventually_gt_at_top 0).mono $ λ x hx, (log_rpow hx _).symm - ... =o[at_top] (λ x, x ^ r) : is_o_log_id_at_top.comp_tendsto (tendsto_rpow_at_top hr) - -lemma is_o_log_rpow_rpow_at_top {s : ℝ} (r : ℝ) (hs : 0 < s) : - (λ x, log x ^ r) =o[at_top] (λ x, x ^ s) := -let r' := max r 1 in -have hr : 0 < r', from lt_max_iff.2 $ or.inr one_pos, -have H : 0 < s / r', from div_pos hs hr, -calc (λ x, log x ^ r) =O[at_top] (λ x, log x ^ r') : - is_O.of_bound 1 $ (tendsto_log_at_top.eventually_ge_at_top 1).mono $ λ x hx, - have hx₀ : 0 ≤ log x, from zero_le_one.trans hx, - by simp [norm_eq_abs, abs_rpow_of_nonneg, abs_rpow_of_nonneg hx₀, - rpow_le_rpow_of_exponent_le (hx.trans (le_abs_self _))] - ... =o[at_top] (λ x, (x ^ (s / r')) ^ r') : - (is_o_log_rpow_at_top H).rpow hr $ (tendsto_rpow_at_top H).eventually $ eventually_ge_at_top 0 - ... =ᶠ[at_top] (λ x, x ^ s) : - (eventually_ge_at_top 0).mono $ λ x hx, by simp only [← rpow_mul hx, div_mul_cancel _ hr.ne'] - -lemma is_o_abs_log_rpow_rpow_nhds_zero {s : ℝ} (r : ℝ) (hs : s < 0) : - (λ x, |log x| ^ r) =o[𝓝[>] 0] (λ x, x ^ s) := -((is_o_log_rpow_rpow_at_top r (neg_pos.2 hs)).comp_tendsto tendsto_inv_zero_at_top).congr' - (mem_of_superset (Icc_mem_nhds_within_Ioi $ set.left_mem_Ico.2 one_pos) $ - λ x hx, by simp [abs_of_nonpos, log_nonpos hx.1 hx.2]) - (eventually_mem_nhds_within.mono $ λ x hx, - by rw [function.comp_app, inv_rpow hx.out.le, rpow_neg hx.out.le, inv_inv]) - -lemma is_o_log_rpow_nhds_zero {r : ℝ} (hr : r < 0) : log =o[𝓝[>] 0] (λ x, x ^ r) := -(is_o_abs_log_rpow_rpow_nhds_zero 1 hr).neg_left.congr' - (mem_of_superset (Icc_mem_nhds_within_Ioi $ set.left_mem_Ico.2 one_pos) $ - λ x hx, by simp [abs_of_nonpos (log_nonpos hx.1 hx.2)]) - eventually_eq.rfl - -lemma tendsto_log_div_rpow_nhds_zero {r : ℝ} (hr : r < 0) : - tendsto (λ x, log x / x ^ r) (𝓝[>] 0) (𝓝 0) := -(is_o_log_rpow_nhds_zero hr).tendsto_div_nhds_zero - -lemma tendsto_log_mul_rpow_nhds_zero {r : ℝ} (hr : 0 < r) : - tendsto (λ x, log x * x ^ r) (𝓝[>] 0) (𝓝 0) := -(tendsto_log_div_rpow_nhds_zero $ neg_lt_zero.2 hr).congr' $ - eventually_mem_nhds_within.mono $ λ x hx, by rw [rpow_neg hx.out.le, div_inv_eq_mul] - -end limits - -namespace complex - -/-- See also `continuous_at_cpow` and `complex.continuous_at_cpow_of_re_pos`. -/ -lemma continuous_at_cpow_zero_of_re_pos {z : ℂ} (hz : 0 < z.re) : - continuous_at (λ x : ℂ × ℂ, x.1 ^ x.2) (0, z) := -begin - have hz₀ : z ≠ 0, from ne_of_apply_ne re hz.ne', - rw [continuous_at, zero_cpow hz₀, tendsto_zero_iff_norm_tendsto_zero], - refine squeeze_zero (λ _, norm_nonneg _) (λ _, abs_cpow_le _ _) _, - simp only [div_eq_mul_inv, ← real.exp_neg], - refine tendsto.zero_mul_is_bounded_under_le _ _, - { convert (continuous_fst.norm.tendsto _).rpow ((continuous_re.comp continuous_snd).tendsto _) _; - simp [hz, real.zero_rpow hz.ne'] }, - { simp only [(∘), real.norm_eq_abs, abs_of_pos (real.exp_pos _)], - rcases exists_gt (|im z|) with ⟨C, hC⟩, - refine ⟨real.exp (π * C), eventually_map.2 _⟩, - refine (((continuous_im.comp continuous_snd).abs.tendsto (_, z)).eventually - (gt_mem_nhds hC)).mono (λ z hz, real.exp_le_exp.2 $ (neg_le_abs_self _).trans _), - rw _root_.abs_mul, - exact mul_le_mul (abs_le.2 ⟨(neg_pi_lt_arg _).le, arg_le_pi _⟩) hz.le - (_root_.abs_nonneg _) real.pi_pos.le } -end - -/-- See also `continuous_at_cpow` for a version that assumes `p.1 ≠ 0` but makes no -assumptions about `p.2`. -/ -lemma continuous_at_cpow_of_re_pos {p : ℂ × ℂ} (h₁ : 0 ≤ p.1.re ∨ p.1.im ≠ 0) (h₂ : 0 < p.2.re) : - continuous_at (λ x : ℂ × ℂ, x.1 ^ x.2) p := -begin - cases p with z w, - rw [← not_lt_zero_iff, lt_iff_le_and_ne, not_and_distrib, ne.def, not_not, not_le_zero_iff] at h₁, - rcases h₁ with h₁|(rfl : z = 0), - exacts [continuous_at_cpow h₁, continuous_at_cpow_zero_of_re_pos h₂] -end - -/-- See also `continuous_at_cpow_const` for a version that assumes `z ≠ 0` but makes no -assumptions about `w`. -/ -lemma continuous_at_cpow_const_of_re_pos {z w : ℂ} (hz : 0 ≤ re z ∨ im z ≠ 0) (hw : 0 < re w) : - continuous_at (λ x, x ^ w) z := -tendsto.comp (@continuous_at_cpow_of_re_pos (z, w) hz hw) - (continuous_at_id.prod continuous_at_const) - -/-- Continuity of `(x, y) ↦ x ^ y` as a function on `ℝ × ℂ`. -/ -lemma continuous_at_of_real_cpow (x : ℝ) (y : ℂ) (h : 0 < y.re ∨ x ≠ 0) : - continuous_at (λ p, ↑p.1 ^ p.2 : ℝ × ℂ → ℂ) (x, y) := -begin - rcases lt_trichotomy 0 x with hx | rfl | hx, - { -- x > 0 : easy case - have : continuous_at (λ p, ⟨↑p.1, p.2⟩ : ℝ × ℂ → ℂ × ℂ) (x, y), - from continuous_of_real.continuous_at.prod_map continuous_at_id, - refine (continuous_at_cpow (or.inl _)).comp this, - rwa of_real_re }, - { -- x = 0 : reduce to continuous_at_cpow_zero_of_re_pos - have A : continuous_at (λ p, p.1 ^ p.2 : ℂ × ℂ → ℂ) ⟨↑(0:ℝ), y⟩, - { rw of_real_zero, - apply continuous_at_cpow_zero_of_re_pos, - tauto }, - have B : continuous_at (λ p, ⟨↑p.1, p.2⟩ : ℝ × ℂ → ℂ × ℂ) ⟨0, y⟩, - from continuous_of_real.continuous_at.prod_map continuous_at_id, - exact @continuous_at.comp (ℝ × ℂ) (ℂ × ℂ) ℂ _ _ _ _ (λ p, ⟨↑p.1, p.2⟩) ⟨0, y⟩ A B }, - { -- x < 0 : difficult case - suffices : continuous_at (λ p, (-↑p.1) ^ p.2 * exp (π * I * p.2) : ℝ × ℂ → ℂ) (x, y), - { refine this.congr (eventually_of_mem (prod_mem_nhds (Iio_mem_nhds hx) univ_mem) _), - exact λ p hp, (of_real_cpow_of_nonpos (le_of_lt hp.1) p.2).symm }, - have A : continuous_at (λ p, ⟨-↑p.1, p.2⟩ : ℝ × ℂ → ℂ × ℂ) (x, y), - from continuous_at.prod_map (continuous_of_real.continuous_at.neg) continuous_at_id, - apply continuous_at.mul, - { refine (continuous_at_cpow (or.inl _)).comp A, - rwa [neg_re, of_real_re, neg_pos] }, - { exact (continuous_exp.comp (continuous_const.mul continuous_snd)).continuous_at } }, -end - -lemma continuous_at_of_real_cpow_const (x : ℝ) (y : ℂ) (h : 0 < y.re ∨ x ≠ 0) : - continuous_at (λ a, a ^ y : ℝ → ℂ) x := -@continuous_at.comp _ _ _ _ _ _ _ _ x (continuous_at_of_real_cpow x y h) - (continuous_id.prod_mk continuous_const).continuous_at - -lemma continuous_of_real_cpow_const {y : ℂ} (hs : 0 < y.re) : continuous (λ x, x ^ y : ℝ → ℂ) := -continuous_iff_continuous_at.mpr (λ x, continuous_at_of_real_cpow_const x y (or.inl hs)) - -end complex - -namespace nnreal - -/-- The nonnegative real power function `x^y`, defined for `x : ℝ≥0` and `y : ℝ ` as the -restriction of the real power function. For `x > 0`, it is equal to `exp (y log x)`. For `x = 0`, -one sets `0 ^ 0 = 1` and `0 ^ y = 0` for `y ≠ 0`. -/ -noncomputable def rpow (x : ℝ≥0) (y : ℝ) : ℝ≥0 := -⟨(x : ℝ) ^ y, real.rpow_nonneg_of_nonneg x.2 y⟩ - -noncomputable instance : has_pow ℝ≥0 ℝ := ⟨rpow⟩ - -@[simp] lemma rpow_eq_pow (x : ℝ≥0) (y : ℝ) : rpow x y = x ^ y := rfl - -@[simp, norm_cast] lemma coe_rpow (x : ℝ≥0) (y : ℝ) : ((x ^ y : ℝ≥0) : ℝ) = (x : ℝ) ^ y := rfl - -@[simp] lemma rpow_zero (x : ℝ≥0) : x ^ (0 : ℝ) = 1 := -nnreal.eq $ real.rpow_zero _ - -@[simp] lemma rpow_eq_zero_iff {x : ℝ≥0} {y : ℝ} : x ^ y = 0 ↔ x = 0 ∧ y ≠ 0 := -begin - rw [← nnreal.coe_eq, coe_rpow, ← nnreal.coe_eq_zero], - exact real.rpow_eq_zero_iff_of_nonneg x.2 -end - -@[simp] lemma zero_rpow {x : ℝ} (h : x ≠ 0) : (0 : ℝ≥0) ^ x = 0 := -nnreal.eq $ real.zero_rpow h - -@[simp] lemma rpow_one (x : ℝ≥0) : x ^ (1 : ℝ) = x := -nnreal.eq $ real.rpow_one _ - -@[simp] lemma one_rpow (x : ℝ) : (1 : ℝ≥0) ^ x = 1 := -nnreal.eq $ real.one_rpow _ - -lemma rpow_add {x : ℝ≥0} (hx : x ≠ 0) (y z : ℝ) : x ^ (y + z) = x ^ y * x ^ z := -nnreal.eq $ real.rpow_add (pos_iff_ne_zero.2 hx) _ _ - -lemma rpow_add' (x : ℝ≥0) {y z : ℝ} (h : y + z ≠ 0) : x ^ (y + z) = x ^ y * x ^ z := -nnreal.eq $ real.rpow_add' x.2 h - -lemma rpow_mul (x : ℝ≥0) (y z : ℝ) : x ^ (y * z) = (x ^ y) ^ z := -nnreal.eq $ real.rpow_mul x.2 y z - -lemma rpow_neg (x : ℝ≥0) (y : ℝ) : x ^ -y = (x ^ y)⁻¹ := -nnreal.eq $ real.rpow_neg x.2 _ - -lemma rpow_neg_one (x : ℝ≥0) : x ^ (-1 : ℝ) = x ⁻¹ := -by simp [rpow_neg] - -lemma rpow_sub {x : ℝ≥0} (hx : x ≠ 0) (y z : ℝ) : x ^ (y - z) = x ^ y / x ^ z := -nnreal.eq $ real.rpow_sub (pos_iff_ne_zero.2 hx) y z - -lemma rpow_sub' (x : ℝ≥0) {y z : ℝ} (h : y - z ≠ 0) : - x ^ (y - z) = x ^ y / x ^ z := -nnreal.eq $ real.rpow_sub' x.2 h - -lemma rpow_inv_rpow_self {y : ℝ} (hy : y ≠ 0) (x : ℝ≥0) : (x ^ y) ^ (1 / y) = x := -by field_simp [← rpow_mul] - -lemma rpow_self_rpow_inv {y : ℝ} (hy : y ≠ 0) (x : ℝ≥0) : (x ^ (1 / y)) ^ y = x := -by field_simp [← rpow_mul] - -lemma inv_rpow (x : ℝ≥0) (y : ℝ) : (x⁻¹) ^ y = (x ^ y)⁻¹ := -nnreal.eq $ real.inv_rpow x.2 y - -lemma div_rpow (x y : ℝ≥0) (z : ℝ) : (x / y) ^ z = x ^ z / y ^ z := -nnreal.eq $ real.div_rpow x.2 y.2 z - -lemma sqrt_eq_rpow (x : ℝ≥0) : sqrt x = x ^ (1/(2:ℝ)) := -begin - refine nnreal.eq _, - push_cast, - exact real.sqrt_eq_rpow x.1, -end - -@[simp, norm_cast] lemma rpow_nat_cast (x : ℝ≥0) (n : ℕ) : x ^ (n : ℝ) = x ^ n := -nnreal.eq $ by simpa only [coe_rpow, coe_pow] using real.rpow_nat_cast x n - -@[simp] lemma rpow_two (x : ℝ≥0) : x ^ (2 : ℝ) = x ^ 2 := -by { rw ← rpow_nat_cast, simp only [nat.cast_bit0, nat.cast_one] } - -lemma mul_rpow {x y : ℝ≥0} {z : ℝ} : (x*y)^z = x^z * y^z := -nnreal.eq $ real.mul_rpow x.2 y.2 - -lemma rpow_le_rpow {x y : ℝ≥0} {z: ℝ} (h₁ : x ≤ y) (h₂ : 0 ≤ z) : x^z ≤ y^z := -real.rpow_le_rpow x.2 h₁ h₂ - -lemma rpow_lt_rpow {x y : ℝ≥0} {z: ℝ} (h₁ : x < y) (h₂ : 0 < z) : x^z < y^z := -real.rpow_lt_rpow x.2 h₁ h₂ - -lemma rpow_lt_rpow_iff {x y : ℝ≥0} {z : ℝ} (hz : 0 < z) : x ^ z < y ^ z ↔ x < y := -real.rpow_lt_rpow_iff x.2 y.2 hz - -lemma rpow_le_rpow_iff {x y : ℝ≥0} {z : ℝ} (hz : 0 < z) : x ^ z ≤ y ^ z ↔ x ≤ y := -real.rpow_le_rpow_iff x.2 y.2 hz - -lemma le_rpow_one_div_iff {x y : ℝ≥0} {z : ℝ} (hz : 0 < z) : x ≤ y ^ (1 / z) ↔ x ^ z ≤ y := -by rw [← rpow_le_rpow_iff hz, rpow_self_rpow_inv hz.ne'] - -lemma rpow_one_div_le_iff {x y : ℝ≥0} {z : ℝ} (hz : 0 < z) : x ^ (1 / z) ≤ y ↔ x ≤ y ^ z := -by rw [← rpow_le_rpow_iff hz, rpow_self_rpow_inv hz.ne'] - -lemma rpow_lt_rpow_of_exponent_lt {x : ℝ≥0} {y z : ℝ} (hx : 1 < x) (hyz : y < z) : x^y < x^z := -real.rpow_lt_rpow_of_exponent_lt hx hyz - -lemma rpow_le_rpow_of_exponent_le {x : ℝ≥0} {y z : ℝ} (hx : 1 ≤ x) (hyz : y ≤ z) : x^y ≤ x^z := -real.rpow_le_rpow_of_exponent_le hx hyz - -lemma rpow_lt_rpow_of_exponent_gt {x : ℝ≥0} {y z : ℝ} (hx0 : 0 < x) (hx1 : x < 1) (hyz : z < y) : - x^y < x^z := -real.rpow_lt_rpow_of_exponent_gt hx0 hx1 hyz - -lemma rpow_le_rpow_of_exponent_ge {x : ℝ≥0} {y z : ℝ} (hx0 : 0 < x) (hx1 : x ≤ 1) (hyz : z ≤ y) : - x^y ≤ x^z := -real.rpow_le_rpow_of_exponent_ge hx0 hx1 hyz - -lemma rpow_pos {p : ℝ} {x : ℝ≥0} (hx_pos : 0 < x) : 0 < x^p := -begin - have rpow_pos_of_nonneg : ∀ {p : ℝ}, 0 < p → 0 < x^p, - { intros p hp_pos, - rw ←zero_rpow hp_pos.ne', - exact rpow_lt_rpow hx_pos hp_pos }, - rcases lt_trichotomy 0 p with hp_pos|rfl|hp_neg, - { exact rpow_pos_of_nonneg hp_pos }, - { simp only [zero_lt_one, rpow_zero] }, - { rw [←neg_neg p, rpow_neg, inv_pos], - exact rpow_pos_of_nonneg (neg_pos.mpr hp_neg) }, -end - -lemma rpow_lt_one {x : ℝ≥0} {z : ℝ} (hx1 : x < 1) (hz : 0 < z) : x^z < 1 := -real.rpow_lt_one (coe_nonneg x) hx1 hz - -lemma rpow_le_one {x : ℝ≥0} {z : ℝ} (hx2 : x ≤ 1) (hz : 0 ≤ z) : x^z ≤ 1 := -real.rpow_le_one x.2 hx2 hz - -lemma rpow_lt_one_of_one_lt_of_neg {x : ℝ≥0} {z : ℝ} (hx : 1 < x) (hz : z < 0) : x^z < 1 := -real.rpow_lt_one_of_one_lt_of_neg hx hz - -lemma rpow_le_one_of_one_le_of_nonpos {x : ℝ≥0} {z : ℝ} (hx : 1 ≤ x) (hz : z ≤ 0) : x^z ≤ 1 := -real.rpow_le_one_of_one_le_of_nonpos hx hz - -lemma one_lt_rpow {x : ℝ≥0} {z : ℝ} (hx : 1 < x) (hz : 0 < z) : 1 < x^z := -real.one_lt_rpow hx hz - -lemma one_le_rpow {x : ℝ≥0} {z : ℝ} (h : 1 ≤ x) (h₁ : 0 ≤ z) : 1 ≤ x^z := -real.one_le_rpow h h₁ - -lemma one_lt_rpow_of_pos_of_lt_one_of_neg {x : ℝ≥0} {z : ℝ} (hx1 : 0 < x) (hx2 : x < 1) - (hz : z < 0) : 1 < x^z := -real.one_lt_rpow_of_pos_of_lt_one_of_neg hx1 hx2 hz - -lemma one_le_rpow_of_pos_of_le_one_of_nonpos {x : ℝ≥0} {z : ℝ} (hx1 : 0 < x) (hx2 : x ≤ 1) - (hz : z ≤ 0) : 1 ≤ x^z := -real.one_le_rpow_of_pos_of_le_one_of_nonpos hx1 hx2 hz - -lemma rpow_le_self_of_le_one {x : ℝ≥0} {z : ℝ} (hx : x ≤ 1) (h_one_le : 1 ≤ z) : x ^ z ≤ x := -begin - rcases eq_bot_or_bot_lt x with rfl | (h : 0 < x), - { have : z ≠ 0 := by linarith, - simp [this] }, - nth_rewrite 1 ←nnreal.rpow_one x, - exact nnreal.rpow_le_rpow_of_exponent_ge h hx h_one_le, -end - -lemma rpow_left_injective {x : ℝ} (hx : x ≠ 0) : function.injective (λ y : ℝ≥0, y^x) := -λ y z hyz, by simpa only [rpow_inv_rpow_self hx] using congr_arg (λ y, y ^ (1 / x)) hyz - -lemma rpow_eq_rpow_iff {x y : ℝ≥0} {z : ℝ} (hz : z ≠ 0) : x ^ z = y ^ z ↔ x = y := -(rpow_left_injective hz).eq_iff - -lemma rpow_left_surjective {x : ℝ} (hx : x ≠ 0) : function.surjective (λ y : ℝ≥0, y^x) := -λ y, ⟨y ^ x⁻¹, by simp_rw [←rpow_mul, _root_.inv_mul_cancel hx, rpow_one]⟩ - -lemma rpow_left_bijective {x : ℝ} (hx : x ≠ 0) : function.bijective (λ y : ℝ≥0, y^x) := -⟨rpow_left_injective hx, rpow_left_surjective hx⟩ - -lemma eq_rpow_one_div_iff {x y : ℝ≥0} {z : ℝ} (hz : z ≠ 0) : x = y ^ (1 / z) ↔ x ^ z = y := -by rw [← rpow_eq_rpow_iff hz, rpow_self_rpow_inv hz] - -lemma rpow_one_div_eq_iff {x y : ℝ≥0} {z : ℝ} (hz : z ≠ 0) : x ^ (1 / z) = y ↔ x = y ^ z := -by rw [← rpow_eq_rpow_iff hz, rpow_self_rpow_inv hz] - -lemma pow_nat_rpow_nat_inv (x : ℝ≥0) {n : ℕ} (hn : n ≠ 0) : - (x ^ n) ^ (n⁻¹ : ℝ) = x := -by { rw [← nnreal.coe_eq, coe_rpow, nnreal.coe_pow], exact real.pow_nat_rpow_nat_inv x.2 hn } - -lemma rpow_nat_inv_pow_nat (x : ℝ≥0) {n : ℕ} (hn : n ≠ 0) : - (x ^ (n⁻¹ : ℝ)) ^ n = x := -by { rw [← nnreal.coe_eq, nnreal.coe_pow, coe_rpow], exact real.rpow_nat_inv_pow_nat x.2 hn } - -lemma continuous_at_rpow {x : ℝ≥0} {y : ℝ} (h : x ≠ 0 ∨ 0 < y) : - continuous_at (λp:ℝ≥0×ℝ, p.1^p.2) (x, y) := -begin - have : (λp:ℝ≥0×ℝ, p.1^p.2) = real.to_nnreal ∘ (λp:ℝ×ℝ, p.1^p.2) ∘ (λp:ℝ≥0 × ℝ, (p.1.1, p.2)), - { ext p, - rw [coe_rpow, real.coe_to_nnreal _ (real.rpow_nonneg_of_nonneg p.1.2 _)], - refl }, - rw this, - refine continuous_real_to_nnreal.continuous_at.comp (continuous_at.comp _ _), - { apply real.continuous_at_rpow, - simp only [ne.def] at h, - rw ← (nnreal.coe_eq_zero x) at h, - exact h }, - { exact ((continuous_subtype_val.comp continuous_fst).prod_mk continuous_snd).continuous_at } -end - -lemma _root_.real.to_nnreal_rpow_of_nonneg {x y : ℝ} (hx : 0 ≤ x) : - real.to_nnreal (x ^ y) = (real.to_nnreal x) ^ y := -begin - nth_rewrite 0 ← real.coe_to_nnreal x hx, - rw [←nnreal.coe_rpow, real.to_nnreal_coe], -end - -lemma eventually_pow_one_div_le (x : ℝ≥0) {y : ℝ≥0} (hy : 1 < y) : - ∀ᶠ (n : ℕ) in at_top, x ^ (1 / n : ℝ) ≤ y := -begin - obtain ⟨m, hm⟩ := add_one_pow_unbounded_of_pos x (tsub_pos_of_lt hy), - rw [tsub_add_cancel_of_le hy.le] at hm, - refine eventually_at_top.2 ⟨m + 1, λ n hn, _⟩, - simpa only [nnreal.rpow_one_div_le_iff (nat.cast_pos.2 $ m.succ_pos.trans_le hn), - nnreal.rpow_nat_cast] using hm.le.trans (pow_le_pow hy.le (m.le_succ.trans hn)), -end - -end nnreal - -namespace real -variables {n : ℕ} - -lemma exists_rat_pow_btwn_rat_aux (hn : n ≠ 0) (x y : ℝ) (h : x < y) (hy : 0 < y) : - ∃ q : ℚ, 0 < q ∧ x < q^n ∧ ↑q^n < y := -begin - have hn' : 0 < (n : ℝ) := by exact_mod_cast hn.bot_lt, - obtain ⟨q, hxq, hqy⟩ := exists_rat_btwn (rpow_lt_rpow (le_max_left 0 x) (max_lt hy h) $ - inv_pos.mpr hn'), - have := rpow_nonneg_of_nonneg (le_max_left 0 x) n⁻¹, - have hq := this.trans_lt hxq, - replace hxq := rpow_lt_rpow this hxq hn', - replace hqy := rpow_lt_rpow hq.le hqy hn', - rw [rpow_nat_cast, rpow_nat_cast, rpow_nat_inv_pow_nat _ hn] at hxq hqy, - exact ⟨q, by exact_mod_cast hq, (le_max_right _ _).trans_lt hxq, hqy⟩, - { exact le_max_left _ _ }, - { exact hy.le } -end - -lemma exists_rat_pow_btwn_rat (hn : n ≠ 0) {x y : ℚ} (h : x < y) (hy : 0 < y) : - ∃ q : ℚ, 0 < q ∧ x < q^n ∧ q^n < y := -by apply_mod_cast exists_rat_pow_btwn_rat_aux hn x y; assumption - -/-- There is a rational power between any two positive elements of an archimedean ordered field. -/ -lemma exists_rat_pow_btwn {α : Type*} [linear_ordered_field α] [archimedean α] (hn : n ≠ 0) - {x y : α} (h : x < y) (hy : 0 < y) : ∃ q : ℚ, 0 < q ∧ x < q^n ∧ (q^n : α) < y := -begin - obtain ⟨q₂, hx₂, hy₂⟩ := exists_rat_btwn (max_lt h hy), - obtain ⟨q₁, hx₁, hq₁₂⟩ := exists_rat_btwn hx₂, - have : (0 : α) < q₂ := (le_max_right _ _).trans_lt hx₂, - norm_cast at hq₁₂ this, - obtain ⟨q, hq, hq₁, hq₂⟩ := exists_rat_pow_btwn_rat hn hq₁₂ this, - refine ⟨q, hq, (le_max_left _ _).trans_lt $ hx₁.trans _, hy₂.trans' _⟩; assumption_mod_cast, -end - -end real - -open filter - -lemma filter.tendsto.nnrpow {α : Type*} {f : filter α} {u : α → ℝ≥0} {v : α → ℝ} {x : ℝ≥0} {y : ℝ} - (hx : tendsto u f (𝓝 x)) (hy : tendsto v f (𝓝 y)) (h : x ≠ 0 ∨ 0 < y) : - tendsto (λ a, (u a) ^ (v a)) f (𝓝 (x ^ y)) := -tendsto.comp (nnreal.continuous_at_rpow h) (hx.prod_mk_nhds hy) - -namespace nnreal - -lemma continuous_at_rpow_const {x : ℝ≥0} {y : ℝ} (h : x ≠ 0 ∨ 0 ≤ y) : - continuous_at (λ z, z^y) x := -h.elim (λ h, tendsto_id.nnrpow tendsto_const_nhds (or.inl h)) $ - λ h, h.eq_or_lt.elim - (λ h, h ▸ by simp only [rpow_zero, continuous_at_const]) - (λ h, tendsto_id.nnrpow tendsto_const_nhds (or.inr h)) - -lemma continuous_rpow_const {y : ℝ} (h : 0 ≤ y) : - continuous (λ x : ℝ≥0, x^y) := -continuous_iff_continuous_at.2 $ λ x, continuous_at_rpow_const (or.inr h) - -theorem tendsto_rpow_at_top {y : ℝ} (hy : 0 < y) : - tendsto (λ (x : ℝ≥0), x ^ y) at_top at_top := -begin - rw filter.tendsto_at_top_at_top, - intros b, - obtain ⟨c, hc⟩ := tendsto_at_top_at_top.mp (tendsto_rpow_at_top hy) b, - use c.to_nnreal, - intros a ha, - exact_mod_cast hc a (real.to_nnreal_le_iff_le_coe.mp ha), -end - -end nnreal - -namespace ennreal - -/-- The real power function `x^y` on extended nonnegative reals, defined for `x : ℝ≥0∞` and -`y : ℝ` as the restriction of the real power function if `0 < x < ⊤`, and with the natural values -for `0` and `⊤` (i.e., `0 ^ x = 0` for `x > 0`, `1` for `x = 0` and `⊤` for `x < 0`, and -`⊤ ^ x = 1 / 0 ^ x`). -/ -noncomputable def rpow : ℝ≥0∞ → ℝ → ℝ≥0∞ -| (some x) y := if x = 0 ∧ y < 0 then ⊤ else (x ^ y : ℝ≥0) -| none y := if 0 < y then ⊤ else if y = 0 then 1 else 0 - -noncomputable instance : has_pow ℝ≥0∞ ℝ := ⟨rpow⟩ - -@[simp] lemma rpow_eq_pow (x : ℝ≥0∞) (y : ℝ) : rpow x y = x ^ y := rfl - -@[simp] lemma rpow_zero {x : ℝ≥0∞} : x ^ (0 : ℝ) = 1 := -by cases x; { dsimp only [(^), rpow], simp [lt_irrefl] } - -lemma top_rpow_def (y : ℝ) : (⊤ : ℝ≥0∞) ^ y = if 0 < y then ⊤ else if y = 0 then 1 else 0 := -rfl - -@[simp] lemma top_rpow_of_pos {y : ℝ} (h : 0 < y) : (⊤ : ℝ≥0∞) ^ y = ⊤ := -by simp [top_rpow_def, h] - -@[simp] lemma top_rpow_of_neg {y : ℝ} (h : y < 0) : (⊤ : ℝ≥0∞) ^ y = 0 := -by simp [top_rpow_def, asymm h, ne_of_lt h] - -@[simp] lemma zero_rpow_of_pos {y : ℝ} (h : 0 < y) : (0 : ℝ≥0∞) ^ y = 0 := -begin - rw [← ennreal.coe_zero, ← ennreal.some_eq_coe], - dsimp only [(^), rpow], - simp [h, asymm h, ne_of_gt h], -end - -@[simp] lemma zero_rpow_of_neg {y : ℝ} (h : y < 0) : (0 : ℝ≥0∞) ^ y = ⊤ := -begin - rw [← ennreal.coe_zero, ← ennreal.some_eq_coe], - dsimp only [(^), rpow], - simp [h, ne_of_gt h], -end - -lemma zero_rpow_def (y : ℝ) : (0 : ℝ≥0∞) ^ y = if 0 < y then 0 else if y = 0 then 1 else ⊤ := -begin - rcases lt_trichotomy 0 y with H|rfl|H, - { simp [H, ne_of_gt, zero_rpow_of_pos, lt_irrefl] }, - { simp [lt_irrefl] }, - { simp [H, asymm H, ne_of_lt, zero_rpow_of_neg] } -end - -@[simp] lemma zero_rpow_mul_self (y : ℝ) : (0 : ℝ≥0∞) ^ y * 0 ^ y = 0 ^ y := -by { rw zero_rpow_def, split_ifs, exacts [zero_mul _, one_mul _, top_mul_top] } - -@[norm_cast] lemma coe_rpow_of_ne_zero {x : ℝ≥0} (h : x ≠ 0) (y : ℝ) : - (x : ℝ≥0∞) ^ y = (x ^ y : ℝ≥0) := -begin - rw [← ennreal.some_eq_coe], - dsimp only [(^), rpow], - simp [h] -end - -@[norm_cast] lemma coe_rpow_of_nonneg (x : ℝ≥0) {y : ℝ} (h : 0 ≤ y) : - (x : ℝ≥0∞) ^ y = (x ^ y : ℝ≥0) := -begin - by_cases hx : x = 0, - { rcases le_iff_eq_or_lt.1 h with H|H, - { simp [hx, H.symm] }, - { simp [hx, zero_rpow_of_pos H, nnreal.zero_rpow (ne_of_gt H)] } }, - { exact coe_rpow_of_ne_zero hx _ } -end - -lemma coe_rpow_def (x : ℝ≥0) (y : ℝ) : - (x : ℝ≥0∞) ^ y = if x = 0 ∧ y < 0 then ⊤ else (x ^ y : ℝ≥0) := rfl - -@[simp] lemma rpow_one (x : ℝ≥0∞) : x ^ (1 : ℝ) = x := -begin - cases x, - { exact dif_pos zero_lt_one }, - { change ite _ _ _ = _, - simp only [nnreal.rpow_one, some_eq_coe, ite_eq_right_iff, top_ne_coe, and_imp], - exact λ _, zero_le_one.not_lt } -end - -@[simp] lemma one_rpow (x : ℝ) : (1 : ℝ≥0∞) ^ x = 1 := -by { rw [← coe_one, coe_rpow_of_ne_zero one_ne_zero], simp } - -@[simp] lemma rpow_eq_zero_iff {x : ℝ≥0∞} {y : ℝ} : - x ^ y = 0 ↔ (x = 0 ∧ 0 < y) ∨ (x = ⊤ ∧ y < 0) := -begin - cases x, - { rcases lt_trichotomy y 0 with H|H|H; - simp [H, top_rpow_of_neg, top_rpow_of_pos, le_of_lt] }, - { by_cases h : x = 0, - { rcases lt_trichotomy y 0 with H|H|H; - simp [h, H, zero_rpow_of_neg, zero_rpow_of_pos, le_of_lt] }, - { simp [coe_rpow_of_ne_zero h, h] } } -end - -@[simp] lemma rpow_eq_top_iff {x : ℝ≥0∞} {y : ℝ} : - x ^ y = ⊤ ↔ (x = 0 ∧ y < 0) ∨ (x = ⊤ ∧ 0 < y) := -begin - cases x, - { rcases lt_trichotomy y 0 with H|H|H; - simp [H, top_rpow_of_neg, top_rpow_of_pos, le_of_lt] }, - { by_cases h : x = 0, - { rcases lt_trichotomy y 0 with H|H|H; - simp [h, H, zero_rpow_of_neg, zero_rpow_of_pos, le_of_lt] }, - { simp [coe_rpow_of_ne_zero h, h] } } -end - -lemma rpow_eq_top_iff_of_pos {x : ℝ≥0∞} {y : ℝ} (hy : 0 < y) : x ^ y = ⊤ ↔ x = ⊤ := -by simp [rpow_eq_top_iff, hy, asymm hy] - -lemma rpow_eq_top_of_nonneg (x : ℝ≥0∞) {y : ℝ} (hy0 : 0 ≤ y) : x ^ y = ⊤ → x = ⊤ := -begin - rw ennreal.rpow_eq_top_iff, - intro h, - cases h, - { exfalso, rw lt_iff_not_ge at h, exact h.right hy0, }, - { exact h.left, }, -end - -lemma rpow_ne_top_of_nonneg {x : ℝ≥0∞} {y : ℝ} (hy0 : 0 ≤ y) (h : x ≠ ⊤) : x ^ y ≠ ⊤ := -mt (ennreal.rpow_eq_top_of_nonneg x hy0) h - -lemma rpow_lt_top_of_nonneg {x : ℝ≥0∞} {y : ℝ} (hy0 : 0 ≤ y) (h : x ≠ ⊤) : x ^ y < ⊤ := -lt_top_iff_ne_top.mpr (ennreal.rpow_ne_top_of_nonneg hy0 h) - -lemma rpow_add {x : ℝ≥0∞} (y z : ℝ) (hx : x ≠ 0) (h'x : x ≠ ⊤) : x ^ (y + z) = x ^ y * x ^ z := -begin - cases x, { exact (h'x rfl).elim }, - have : x ≠ 0 := λ h, by simpa [h] using hx, - simp [coe_rpow_of_ne_zero this, nnreal.rpow_add this] -end - -lemma rpow_neg (x : ℝ≥0∞) (y : ℝ) : x ^ -y = (x ^ y)⁻¹ := -begin - cases x, - { rcases lt_trichotomy y 0 with H|H|H; - simp [top_rpow_of_pos, top_rpow_of_neg, H, neg_pos.mpr] }, - { by_cases h : x = 0, - { rcases lt_trichotomy y 0 with H|H|H; - simp [h, zero_rpow_of_pos, zero_rpow_of_neg, H, neg_pos.mpr] }, - { have A : x ^ y ≠ 0, by simp [h], - simp [coe_rpow_of_ne_zero h, ← coe_inv A, nnreal.rpow_neg] } } -end - -lemma rpow_sub {x : ℝ≥0∞} (y z : ℝ) (hx : x ≠ 0) (h'x : x ≠ ⊤) : x ^ (y - z) = x ^ y / x ^ z := -by rw [sub_eq_add_neg, rpow_add _ _ hx h'x, rpow_neg, div_eq_mul_inv] - -lemma rpow_neg_one (x : ℝ≥0∞) : x ^ (-1 : ℝ) = x ⁻¹ := -by simp [rpow_neg] - -lemma rpow_mul (x : ℝ≥0∞) (y z : ℝ) : x ^ (y * z) = (x ^ y) ^ z := -begin - cases x, - { rcases lt_trichotomy y 0 with Hy|Hy|Hy; - rcases lt_trichotomy z 0 with Hz|Hz|Hz; - simp [Hy, Hz, zero_rpow_of_neg, zero_rpow_of_pos, top_rpow_of_neg, top_rpow_of_pos, - mul_pos_of_neg_of_neg, mul_neg_of_neg_of_pos, mul_neg_of_pos_of_neg] }, - { by_cases h : x = 0, - { rcases lt_trichotomy y 0 with Hy|Hy|Hy; - rcases lt_trichotomy z 0 with Hz|Hz|Hz; - simp [h, Hy, Hz, zero_rpow_of_neg, zero_rpow_of_pos, top_rpow_of_neg, top_rpow_of_pos, - mul_pos_of_neg_of_neg, mul_neg_of_neg_of_pos, mul_neg_of_pos_of_neg] }, - { have : x ^ y ≠ 0, by simp [h], - simp [coe_rpow_of_ne_zero h, coe_rpow_of_ne_zero this, nnreal.rpow_mul] } } -end - -@[simp, norm_cast] lemma rpow_nat_cast (x : ℝ≥0∞) (n : ℕ) : x ^ (n : ℝ) = x ^ n := -begin - cases x, - { cases n; - simp [top_rpow_of_pos (nat.cast_add_one_pos _), top_pow (nat.succ_pos _)] }, - { simp [coe_rpow_of_nonneg _ (nat.cast_nonneg n)] } -end - -@[simp] lemma rpow_two (x : ℝ≥0∞) : x ^ (2 : ℝ) = x ^ 2 := -by { rw ← rpow_nat_cast, simp only [nat.cast_bit0, nat.cast_one] } - -lemma mul_rpow_eq_ite (x y : ℝ≥0∞) (z : ℝ) : - (x * y) ^ z = if (x = 0 ∧ y = ⊤ ∨ x = ⊤ ∧ y = 0) ∧ z < 0 then ⊤ else x ^ z * y ^ z := -begin - rcases eq_or_ne z 0 with rfl|hz, { simp }, - replace hz := hz.lt_or_lt, - wlog hxy : x ≤ y, - { convert this y x z hz (le_of_not_le hxy) using 2; simp only [mul_comm, and_comm, or_comm], }, - rcases eq_or_ne x 0 with rfl|hx0, - { induction y using with_top.rec_top_coe; cases hz with hz hz; simp [*, hz.not_lt] }, - rcases eq_or_ne y 0 with rfl|hy0, { exact (hx0 (bot_unique hxy)).elim }, - induction x using with_top.rec_top_coe, { cases hz with hz hz; simp [hz, top_unique hxy] }, - induction y using with_top.rec_top_coe, { cases hz with hz hz; simp * }, - simp only [*, false_and, and_false, false_or, if_false], - norm_cast at *, - rw [coe_rpow_of_ne_zero (mul_ne_zero hx0 hy0), nnreal.mul_rpow] -end - -lemma mul_rpow_of_ne_top {x y : ℝ≥0∞} (hx : x ≠ ⊤) (hy : y ≠ ⊤) (z : ℝ) : - (x * y) ^ z = x^z * y^z := -by simp [*, mul_rpow_eq_ite] - -@[norm_cast] lemma coe_mul_rpow (x y : ℝ≥0) (z : ℝ) : - ((x : ℝ≥0∞) * y) ^ z = x^z * y^z := -mul_rpow_of_ne_top coe_ne_top coe_ne_top z - -lemma mul_rpow_of_ne_zero {x y : ℝ≥0∞} (hx : x ≠ 0) (hy : y ≠ 0) (z : ℝ) : - (x * y) ^ z = x ^ z * y ^ z := -by simp [*, mul_rpow_eq_ite] - -lemma mul_rpow_of_nonneg (x y : ℝ≥0∞) {z : ℝ} (hz : 0 ≤ z) : - (x * y) ^ z = x ^ z * y ^ z := -by simp [hz.not_lt, mul_rpow_eq_ite] - -lemma inv_rpow (x : ℝ≥0∞) (y : ℝ) : (x⁻¹) ^ y = (x ^ y)⁻¹ := -begin - rcases eq_or_ne y 0 with rfl|hy, { simp only [rpow_zero, inv_one] }, - replace hy := hy.lt_or_lt, - rcases eq_or_ne x 0 with rfl|h0, { cases hy; simp * }, - rcases eq_or_ne x ⊤ with rfl|h_top, { cases hy; simp * }, - apply ennreal.eq_inv_of_mul_eq_one_left, - rw [← mul_rpow_of_ne_zero (ennreal.inv_ne_zero.2 h_top) h0, ennreal.inv_mul_cancel h0 h_top, - one_rpow] -end - -lemma div_rpow_of_nonneg (x y : ℝ≥0∞) {z : ℝ} (hz : 0 ≤ z) : - (x / y) ^ z = x ^ z / y ^ z := -by rw [div_eq_mul_inv, mul_rpow_of_nonneg _ _ hz, inv_rpow, div_eq_mul_inv] - -lemma strict_mono_rpow_of_pos {z : ℝ} (h : 0 < z) : strict_mono (λ x : ℝ≥0∞, x ^ z) := -begin - intros x y hxy, - lift x to ℝ≥0 using ne_top_of_lt hxy, - rcases eq_or_ne y ∞ with rfl|hy, - { simp only [top_rpow_of_pos h, coe_rpow_of_nonneg _ h.le, coe_lt_top] }, - { lift y to ℝ≥0 using hy, - simp only [coe_rpow_of_nonneg _ h.le, nnreal.rpow_lt_rpow (coe_lt_coe.1 hxy) h, coe_lt_coe] } -end - -lemma monotone_rpow_of_nonneg {z : ℝ} (h : 0 ≤ z) : monotone (λ x : ℝ≥0∞, x ^ z) := -h.eq_or_lt.elim (λ h0, h0 ▸ by simp only [rpow_zero, monotone_const]) - (λ h0, (strict_mono_rpow_of_pos h0).monotone) - -/-- Bundles `λ x : ℝ≥0∞, x ^ y` into an order isomorphism when `y : ℝ` is positive, -where the inverse is `λ x : ℝ≥0∞, x ^ (1 / y)`. -/ -@[simps apply] def order_iso_rpow (y : ℝ) (hy : 0 < y) : ℝ≥0∞ ≃o ℝ≥0∞ := -(strict_mono_rpow_of_pos hy).order_iso_of_right_inverse (λ x, x ^ y) (λ x, x ^ (1 / y)) - (λ x, by { dsimp, rw [←rpow_mul, one_div_mul_cancel hy.ne.symm, rpow_one] }) - -lemma order_iso_rpow_symm_apply (y : ℝ) (hy : 0 < y) : - (order_iso_rpow y hy).symm = order_iso_rpow (1 / y) (one_div_pos.2 hy) := -by { simp only [order_iso_rpow, one_div_one_div], refl } - -lemma rpow_le_rpow {x y : ℝ≥0∞} {z : ℝ} (h₁ : x ≤ y) (h₂ : 0 ≤ z) : x^z ≤ y^z := -monotone_rpow_of_nonneg h₂ h₁ - -lemma rpow_lt_rpow {x y : ℝ≥0∞} {z : ℝ} (h₁ : x < y) (h₂ : 0 < z) : x^z < y^z := -strict_mono_rpow_of_pos h₂ h₁ - -lemma rpow_le_rpow_iff {x y : ℝ≥0∞} {z : ℝ} (hz : 0 < z) : x ^ z ≤ y ^ z ↔ x ≤ y := -(strict_mono_rpow_of_pos hz).le_iff_le - -lemma rpow_lt_rpow_iff {x y : ℝ≥0∞} {z : ℝ} (hz : 0 < z) : x ^ z < y ^ z ↔ x < y := -(strict_mono_rpow_of_pos hz).lt_iff_lt - -lemma le_rpow_one_div_iff {x y : ℝ≥0∞} {z : ℝ} (hz : 0 < z) : x ≤ y ^ (1 / z) ↔ x ^ z ≤ y := -begin - nth_rewrite 0 ←rpow_one x, - nth_rewrite 0 ←@_root_.mul_inv_cancel _ _ z hz.ne', - rw [rpow_mul, ←one_div, @rpow_le_rpow_iff _ _ (1/z) (by simp [hz])], -end - -lemma lt_rpow_one_div_iff {x y : ℝ≥0∞} {z : ℝ} (hz : 0 < z) : x < y ^ (1 / z) ↔ x ^ z < y := -begin - nth_rewrite 0 ←rpow_one x, - nth_rewrite 0 ←@_root_.mul_inv_cancel _ _ z (ne_of_lt hz).symm, - rw [rpow_mul, ←one_div, @rpow_lt_rpow_iff _ _ (1/z) (by simp [hz])], -end - -lemma rpow_one_div_le_iff {x y : ℝ≥0∞} {z : ℝ} (hz : 0 < z) : x ^ (1 / z) ≤ y ↔ x ≤ y ^ z := -begin - nth_rewrite 0 ← ennreal.rpow_one y, - nth_rewrite 1 ← @_root_.mul_inv_cancel _ _ z hz.ne.symm, - rw [ennreal.rpow_mul, ← one_div, ennreal.rpow_le_rpow_iff (one_div_pos.2 hz)], -end - -lemma rpow_lt_rpow_of_exponent_lt {x : ℝ≥0∞} {y z : ℝ} (hx : 1 < x) (hx' : x ≠ ⊤) (hyz : y < z) : - x^y < x^z := -begin - lift x to ℝ≥0 using hx', - rw [one_lt_coe_iff] at hx, - simp [coe_rpow_of_ne_zero (ne_of_gt (lt_trans zero_lt_one hx)), - nnreal.rpow_lt_rpow_of_exponent_lt hx hyz] -end - -lemma rpow_le_rpow_of_exponent_le {x : ℝ≥0∞} {y z : ℝ} (hx : 1 ≤ x) (hyz : y ≤ z) : x^y ≤ x^z := -begin - cases x, - { rcases lt_trichotomy y 0 with Hy|Hy|Hy; - rcases lt_trichotomy z 0 with Hz|Hz|Hz; - simp [Hy, Hz, top_rpow_of_neg, top_rpow_of_pos, le_refl]; - linarith }, - { simp only [one_le_coe_iff, some_eq_coe] at hx, - simp [coe_rpow_of_ne_zero (ne_of_gt (lt_of_lt_of_le zero_lt_one hx)), - nnreal.rpow_le_rpow_of_exponent_le hx hyz] } -end - -lemma rpow_lt_rpow_of_exponent_gt {x : ℝ≥0∞} {y z : ℝ} (hx0 : 0 < x) (hx1 : x < 1) (hyz : z < y) : - x^y < x^z := -begin - lift x to ℝ≥0 using ne_of_lt (lt_of_lt_of_le hx1 le_top), - simp only [coe_lt_one_iff, coe_pos] at hx0 hx1, - simp [coe_rpow_of_ne_zero (ne_of_gt hx0), nnreal.rpow_lt_rpow_of_exponent_gt hx0 hx1 hyz] -end - -lemma rpow_le_rpow_of_exponent_ge {x : ℝ≥0∞} {y z : ℝ} (hx1 : x ≤ 1) (hyz : z ≤ y) : - x^y ≤ x^z := -begin - lift x to ℝ≥0 using ne_of_lt (lt_of_le_of_lt hx1 coe_lt_top), - by_cases h : x = 0, - { rcases lt_trichotomy y 0 with Hy|Hy|Hy; - rcases lt_trichotomy z 0 with Hz|Hz|Hz; - simp [Hy, Hz, h, zero_rpow_of_neg, zero_rpow_of_pos, le_refl]; - linarith }, - { rw [coe_le_one_iff] at hx1, - simp [coe_rpow_of_ne_zero h, - nnreal.rpow_le_rpow_of_exponent_ge (bot_lt_iff_ne_bot.mpr h) hx1 hyz] } -end - -lemma rpow_le_self_of_le_one {x : ℝ≥0∞} {z : ℝ} (hx : x ≤ 1) (h_one_le : 1 ≤ z) : x ^ z ≤ x := -begin - nth_rewrite 1 ←ennreal.rpow_one x, - exact ennreal.rpow_le_rpow_of_exponent_ge hx h_one_le, -end - -lemma le_rpow_self_of_one_le {x : ℝ≥0∞} {z : ℝ} (hx : 1 ≤ x) (h_one_le : 1 ≤ z) : x ≤ x ^ z := -begin - nth_rewrite 0 ←ennreal.rpow_one x, - exact ennreal.rpow_le_rpow_of_exponent_le hx h_one_le, -end - -lemma rpow_pos_of_nonneg {p : ℝ} {x : ℝ≥0∞} (hx_pos : 0 < x) (hp_nonneg : 0 ≤ p) : 0 < x^p := -begin - by_cases hp_zero : p = 0, - { simp [hp_zero, zero_lt_one], }, - { rw ←ne.def at hp_zero, - have hp_pos := lt_of_le_of_ne hp_nonneg hp_zero.symm, - rw ←zero_rpow_of_pos hp_pos, exact rpow_lt_rpow hx_pos hp_pos, }, -end - -lemma rpow_pos {p : ℝ} {x : ℝ≥0∞} (hx_pos : 0 < x) (hx_ne_top : x ≠ ⊤) : 0 < x^p := -begin - cases lt_or_le 0 p with hp_pos hp_nonpos, - { exact rpow_pos_of_nonneg hx_pos (le_of_lt hp_pos), }, - { rw [←neg_neg p, rpow_neg, ennreal.inv_pos], - exact rpow_ne_top_of_nonneg (right.nonneg_neg_iff.mpr hp_nonpos) hx_ne_top, }, -end - -lemma rpow_lt_one {x : ℝ≥0∞} {z : ℝ} (hx : x < 1) (hz : 0 < z) : x^z < 1 := -begin - lift x to ℝ≥0 using ne_of_lt (lt_of_lt_of_le hx le_top), - simp only [coe_lt_one_iff] at hx, - simp [coe_rpow_of_nonneg _ (le_of_lt hz), nnreal.rpow_lt_one hx hz], -end - -lemma rpow_le_one {x : ℝ≥0∞} {z : ℝ} (hx : x ≤ 1) (hz : 0 ≤ z) : x^z ≤ 1 := -begin - lift x to ℝ≥0 using ne_of_lt (lt_of_le_of_lt hx coe_lt_top), - simp only [coe_le_one_iff] at hx, - simp [coe_rpow_of_nonneg _ hz, nnreal.rpow_le_one hx hz], -end - -lemma rpow_lt_one_of_one_lt_of_neg {x : ℝ≥0∞} {z : ℝ} (hx : 1 < x) (hz : z < 0) : x^z < 1 := -begin - cases x, - { simp [top_rpow_of_neg hz, zero_lt_one] }, - { simp only [some_eq_coe, one_lt_coe_iff] at hx, - simp [coe_rpow_of_ne_zero (ne_of_gt (lt_trans zero_lt_one hx)), - nnreal.rpow_lt_one_of_one_lt_of_neg hx hz] }, -end - -lemma rpow_le_one_of_one_le_of_neg {x : ℝ≥0∞} {z : ℝ} (hx : 1 ≤ x) (hz : z < 0) : x^z ≤ 1 := -begin - cases x, - { simp [top_rpow_of_neg hz, zero_lt_one] }, - { simp only [one_le_coe_iff, some_eq_coe] at hx, - simp [coe_rpow_of_ne_zero (ne_of_gt (lt_of_lt_of_le zero_lt_one hx)), - nnreal.rpow_le_one_of_one_le_of_nonpos hx (le_of_lt hz)] }, -end - -lemma one_lt_rpow {x : ℝ≥0∞} {z : ℝ} (hx : 1 < x) (hz : 0 < z) : 1 < x^z := -begin - cases x, - { simp [top_rpow_of_pos hz] }, - { simp only [some_eq_coe, one_lt_coe_iff] at hx, - simp [coe_rpow_of_nonneg _ (le_of_lt hz), nnreal.one_lt_rpow hx hz] } -end - -lemma one_le_rpow {x : ℝ≥0∞} {z : ℝ} (hx : 1 ≤ x) (hz : 0 < z) : 1 ≤ x^z := -begin - cases x, - { simp [top_rpow_of_pos hz] }, - { simp only [one_le_coe_iff, some_eq_coe] at hx, - simp [coe_rpow_of_nonneg _ (le_of_lt hz), nnreal.one_le_rpow hx (le_of_lt hz)] }, -end - -lemma one_lt_rpow_of_pos_of_lt_one_of_neg {x : ℝ≥0∞} {z : ℝ} (hx1 : 0 < x) (hx2 : x < 1) - (hz : z < 0) : 1 < x^z := -begin - lift x to ℝ≥0 using ne_of_lt (lt_of_lt_of_le hx2 le_top), - simp only [coe_lt_one_iff, coe_pos] at ⊢ hx1 hx2, - simp [coe_rpow_of_ne_zero (ne_of_gt hx1), nnreal.one_lt_rpow_of_pos_of_lt_one_of_neg hx1 hx2 hz], -end - -lemma one_le_rpow_of_pos_of_le_one_of_neg {x : ℝ≥0∞} {z : ℝ} (hx1 : 0 < x) (hx2 : x ≤ 1) - (hz : z < 0) : 1 ≤ x^z := -begin - lift x to ℝ≥0 using ne_of_lt (lt_of_le_of_lt hx2 coe_lt_top), - simp only [coe_le_one_iff, coe_pos] at ⊢ hx1 hx2, - simp [coe_rpow_of_ne_zero (ne_of_gt hx1), - nnreal.one_le_rpow_of_pos_of_le_one_of_nonpos hx1 hx2 (le_of_lt hz)], -end - -lemma to_nnreal_rpow (x : ℝ≥0∞) (z : ℝ) : (x.to_nnreal) ^ z = (x ^ z).to_nnreal := -begin - rcases lt_trichotomy z 0 with H|H|H, - { cases x, { simp [H, ne_of_lt] }, - by_cases hx : x = 0, - { simp [hx, H, ne_of_lt] }, - { simp [coe_rpow_of_ne_zero hx] } }, - { simp [H] }, - { cases x, { simp [H, ne_of_gt] }, - simp [coe_rpow_of_nonneg _ (le_of_lt H)] } -end - -lemma to_real_rpow (x : ℝ≥0∞) (z : ℝ) : (x.to_real) ^ z = (x ^ z).to_real := -by rw [ennreal.to_real, ennreal.to_real, ←nnreal.coe_rpow, ennreal.to_nnreal_rpow] - -lemma of_real_rpow_of_pos {x p : ℝ} (hx_pos : 0 < x) : - ennreal.of_real x ^ p = ennreal.of_real (x ^ p) := -begin - simp_rw ennreal.of_real, - rw [coe_rpow_of_ne_zero, coe_eq_coe, real.to_nnreal_rpow_of_nonneg hx_pos.le], - simp [hx_pos], -end - -lemma of_real_rpow_of_nonneg {x p : ℝ} (hx_nonneg : 0 ≤ x) (hp_nonneg : 0 ≤ p) : - ennreal.of_real x ^ p = ennreal.of_real (x ^ p) := -begin - by_cases hp0 : p = 0, - { simp [hp0], }, - by_cases hx0 : x = 0, - { rw ← ne.def at hp0, - have hp_pos : 0 < p := lt_of_le_of_ne hp_nonneg hp0.symm, - simp [hx0, hp_pos, hp_pos.ne.symm], }, - rw ← ne.def at hx0, - exact of_real_rpow_of_pos (hx_nonneg.lt_of_ne hx0.symm), -end - -lemma rpow_left_injective {x : ℝ} (hx : x ≠ 0) : - function.injective (λ y : ℝ≥0∞, y^x) := -begin - intros y z hyz, - dsimp only at hyz, - rw [←rpow_one y, ←rpow_one z, ←_root_.mul_inv_cancel hx, rpow_mul, rpow_mul, hyz], -end - -lemma rpow_left_surjective {x : ℝ} (hx : x ≠ 0) : - function.surjective (λ y : ℝ≥0∞, y^x) := -λ y, ⟨y ^ x⁻¹, by simp_rw [←rpow_mul, _root_.inv_mul_cancel hx, rpow_one]⟩ - -lemma rpow_left_bijective {x : ℝ} (hx : x ≠ 0) : - function.bijective (λ y : ℝ≥0∞, y^x) := -⟨rpow_left_injective hx, rpow_left_surjective hx⟩ - -theorem tendsto_rpow_at_top {y : ℝ} (hy : 0 < y) : - tendsto (λ (x : ℝ≥0∞), x ^ y) (𝓝 ⊤) (𝓝 ⊤) := -begin - rw tendsto_nhds_top_iff_nnreal, - intros x, - obtain ⟨c, _, hc⟩ := - (at_top_basis_Ioi.tendsto_iff at_top_basis_Ioi).mp (nnreal.tendsto_rpow_at_top hy) x trivial, - have hc' : set.Ioi (↑c) ∈ 𝓝 (⊤ : ℝ≥0∞) := Ioi_mem_nhds coe_lt_top, - refine eventually_of_mem hc' _, - intros a ha, - by_cases ha' : a = ⊤, - { simp [ha', hy] }, - lift a to ℝ≥0 using ha', - change ↑c < ↑a at ha, - rw coe_rpow_of_nonneg _ hy.le, - exact_mod_cast hc a (by exact_mod_cast ha), -end - -lemma eventually_pow_one_div_le {x : ℝ≥0∞} (hx : x ≠ ∞) {y : ℝ≥0∞} (hy : 1 < y) : - ∀ᶠ (n : ℕ) in at_top, x ^ (1 / n : ℝ) ≤ y := -begin - lift x to ℝ≥0 using hx, - by_cases y = ∞, - { exact eventually_of_forall (λ n, h.symm ▸ le_top) }, - { lift y to ℝ≥0 using h, - have := nnreal.eventually_pow_one_div_le x (by exact_mod_cast hy : 1 < y), - refine this.congr (eventually_of_forall $ λ n, _), - rw [coe_rpow_of_nonneg x (by positivity : 0 ≤ (1 / n : ℝ)), coe_le_coe] }, -end - -private lemma continuous_at_rpow_const_of_pos {x : ℝ≥0∞} {y : ℝ} (h : 0 < y) : - continuous_at (λ a : ℝ≥0∞, a ^ y) x := -begin - by_cases hx : x = ⊤, - { rw [hx, continuous_at], - convert tendsto_rpow_at_top h, - simp [h] }, - lift x to ℝ≥0 using hx, - rw continuous_at_coe_iff, - convert continuous_coe.continuous_at.comp - (nnreal.continuous_at_rpow_const (or.inr h.le)) using 1, - ext1 x, - simp [coe_rpow_of_nonneg _ h.le] -end - -@[continuity] -lemma continuous_rpow_const {y : ℝ} : continuous (λ a : ℝ≥0∞, a ^ y) := -begin - apply continuous_iff_continuous_at.2 (λ x, _), - rcases lt_trichotomy 0 y with hy|rfl|hy, - { exact continuous_at_rpow_const_of_pos hy }, - { simp only [rpow_zero], exact continuous_at_const }, - { obtain ⟨z, hz⟩ : ∃ z, y = -z := ⟨-y, (neg_neg _).symm⟩, - have z_pos : 0 < z, by simpa [hz] using hy, - simp_rw [hz, rpow_neg], - exact continuous_inv.continuous_at.comp (continuous_at_rpow_const_of_pos z_pos) } -end - -lemma tendsto_const_mul_rpow_nhds_zero_of_pos {c : ℝ≥0∞} (hc : c ≠ ∞) {y : ℝ} (hy : 0 < y) : - tendsto (λ x : ℝ≥0∞, c * x ^ y) (𝓝 0) (𝓝 0) := -begin - convert ennreal.tendsto.const_mul (ennreal.continuous_rpow_const.tendsto 0) _, - { simp [hy] }, - { exact or.inr hc } -end - -end ennreal - -lemma filter.tendsto.ennrpow_const {α : Type*} {f : filter α} {m : α → ℝ≥0∞} {a : ℝ≥0∞} (r : ℝ) - (hm : tendsto m f (𝓝 a)) : - tendsto (λ x, (m x) ^ r) f (𝓝 (a ^ r)) := -(ennreal.continuous_rpow_const.tendsto a).comp hm - -namespace norm_num -open tactic - -theorem rpow_pos (a b : ℝ) (b' : ℕ) (c : ℝ) (hb : (b':ℝ) = b) (h : a ^ b' = c) : a ^ b = c := -by rw [← h, ← hb, real.rpow_nat_cast] -theorem rpow_neg (a b : ℝ) (b' : ℕ) (c c' : ℝ) - (a0 : 0 ≤ a) (hb : (b':ℝ) = b) (h : a ^ b' = c) (hc : c⁻¹ = c') : a ^ -b = c' := -by rw [← hc, ← h, ← hb, real.rpow_neg a0, real.rpow_nat_cast] - -/-- Evaluate `real.rpow a b` where `a` is a rational numeral and `b` is an integer. -(This cannot go via the generalized version `prove_rpow'` because `rpow_pos` has a side condition; -we do not attempt to evaluate `a ^ b` where `a` and `b` are both negative because it comes -out to some garbage.) -/ -meta def prove_rpow (a b : expr) : tactic (expr × expr) := do - na ← a.to_rat, - ic ← mk_instance_cache `(ℝ), - match match_sign b with - | sum.inl b := do - (ic, a0) ← guard (na ≥ 0) >> prove_nonneg ic a, - nc ← mk_instance_cache `(ℕ), - (ic, nc, b', hb) ← prove_nat_uncast ic nc b, - (ic, c, h) ← prove_pow a na ic b', - cr ← c.to_rat, - (ic, c', hc) ← prove_inv ic c cr, - pure (c', (expr.const ``rpow_neg []).mk_app [a, b, b', c, c', a0, hb, h, hc]) - | sum.inr ff := pure (`(1:ℝ), expr.const ``real.rpow_zero [] a) - | sum.inr tt := do - nc ← mk_instance_cache `(ℕ), - (ic, nc, b', hb) ← prove_nat_uncast ic nc b, - (ic, c, h) ← prove_pow a na ic b', - pure (c, (expr.const ``rpow_pos []).mk_app [a, b, b', c, hb, h]) - end - -/-- Generalized version of `prove_cpow`, `prove_nnrpow`, `prove_ennrpow`. -/ -meta def prove_rpow' (pos neg zero : name) (α β one a b : expr) : tactic (expr × expr) := do - na ← a.to_rat, - icα ← mk_instance_cache α, - icβ ← mk_instance_cache β, - match match_sign b with - | sum.inl b := do - nc ← mk_instance_cache `(ℕ), - (icβ, nc, b', hb) ← prove_nat_uncast icβ nc b, - (icα, c, h) ← prove_pow a na icα b', - cr ← c.to_rat, - (icα, c', hc) ← prove_inv icα c cr, - pure (c', (expr.const neg []).mk_app [a, b, b', c, c', hb, h, hc]) - | sum.inr ff := pure (one, expr.const zero [] a) - | sum.inr tt := do - nc ← mk_instance_cache `(ℕ), - (icβ, nc, b', hb) ← prove_nat_uncast icβ nc b, - (icα, c, h) ← prove_pow a na icα b', - pure (c, (expr.const pos []).mk_app [a, b, b', c, hb, h]) - end - -open_locale nnreal ennreal - -theorem cpow_pos (a b : ℂ) (b' : ℕ) (c : ℂ) (hb : b = b') (h : a ^ b' = c) : a ^ b = c := -by rw [← h, hb, complex.cpow_nat_cast] -theorem cpow_neg (a b : ℂ) (b' : ℕ) (c c' : ℂ) - (hb : b = b') (h : a ^ b' = c) (hc : c⁻¹ = c') : a ^ -b = c' := -by rw [← hc, ← h, hb, complex.cpow_neg, complex.cpow_nat_cast] - -theorem nnrpow_pos (a : ℝ≥0) (b : ℝ) (b' : ℕ) (c : ℝ≥0) - (hb : b = b') (h : a ^ b' = c) : a ^ b = c := -by rw [← h, hb, nnreal.rpow_nat_cast] -theorem nnrpow_neg (a : ℝ≥0) (b : ℝ) (b' : ℕ) (c c' : ℝ≥0) - (hb : b = b') (h : a ^ b' = c) (hc : c⁻¹ = c') : a ^ -b = c' := -by rw [← hc, ← h, hb, nnreal.rpow_neg, nnreal.rpow_nat_cast] - -theorem ennrpow_pos (a : ℝ≥0∞) (b : ℝ) (b' : ℕ) (c : ℝ≥0∞) - (hb : b = b') (h : a ^ b' = c) : a ^ b = c := -by rw [← h, hb, ennreal.rpow_nat_cast] -theorem ennrpow_neg (a : ℝ≥0∞) (b : ℝ) (b' : ℕ) (c c' : ℝ≥0∞) - (hb : b = b') (h : a ^ b' = c) (hc : c⁻¹ = c') : a ^ -b = c' := -by rw [← hc, ← h, hb, ennreal.rpow_neg, ennreal.rpow_nat_cast] - -/-- Evaluate `complex.cpow a b` where `a` is a rational numeral and `b` is an integer. -/ -meta def prove_cpow : expr → expr → tactic (expr × expr) := -prove_rpow' ``cpow_pos ``cpow_neg ``complex.cpow_zero `(ℂ) `(ℂ) `(1:ℂ) - -/-- Evaluate `nnreal.rpow a b` where `a` is a rational numeral and `b` is an integer. -/ -meta def prove_nnrpow : expr → expr → tactic (expr × expr) := -prove_rpow' ``nnrpow_pos ``nnrpow_neg ``nnreal.rpow_zero `(ℝ≥0) `(ℝ) `(1:ℝ≥0) - -/-- Evaluate `ennreal.rpow a b` where `a` is a rational numeral and `b` is an integer. -/ -meta def prove_ennrpow : expr → expr → tactic (expr × expr) := -prove_rpow' ``ennrpow_pos ``ennrpow_neg ``ennreal.rpow_zero `(ℝ≥0∞) `(ℝ) `(1:ℝ≥0∞) - -/-- Evaluates expressions of the form `rpow a b`, `cpow a b` and `a ^ b` in the special case where -`b` is an integer and `a` is a positive rational (so it's really just a rational power). -/ -@[norm_num] meta def eval_rpow_cpow : expr → tactic (expr × expr) -| `(@has_pow.pow _ _ real.has_pow %%a %%b) := b.to_int >> prove_rpow a b -| `(real.rpow %%a %%b) := b.to_int >> prove_rpow a b -| `(@has_pow.pow _ _ complex.has_pow %%a %%b) := b.to_int >> prove_cpow a b -| `(complex.cpow %%a %%b) := b.to_int >> prove_cpow a b -| `(@has_pow.pow _ _ nnreal.real.has_pow %%a %%b) := b.to_int >> prove_nnrpow a b -| `(nnreal.rpow %%a %%b) := b.to_int >> prove_nnrpow a b -| `(@has_pow.pow _ _ ennreal.real.has_pow %%a %%b) := b.to_int >> prove_ennrpow a b -| `(ennreal.rpow %%a %%b) := b.to_int >> prove_ennrpow a b -| _ := tactic.failed - -end norm_num - -namespace tactic -namespace positivity - -/-- Auxiliary definition for the `positivity` tactic to handle real powers of reals. -/ -meta def prove_rpow (a b : expr) : tactic strictness := -do - strictness_a ← core a, - match strictness_a with - | nonnegative p := nonnegative <$> mk_app ``real.rpow_nonneg_of_nonneg [p, b] - | positive p := positive <$> mk_app ``real.rpow_pos_of_pos [p, b] - | _ := failed - end - -private lemma nnrpow_pos {a : ℝ≥0} (ha : 0 < a) (b : ℝ) : 0 < a ^ b := nnreal.rpow_pos ha - -/-- Auxiliary definition for the `positivity` tactic to handle real powers of nonnegative reals. -/ -meta def prove_nnrpow (a b : expr) : tactic strictness := -do - strictness_a ← core a, - match strictness_a with - | positive p := positive <$> mk_app ``nnrpow_pos [p, b] - | _ := failed -- We already know `0 ≤ x` for all `x : ℝ≥0` - end - -private lemma ennrpow_pos {a : ℝ≥0∞} {b : ℝ} (ha : 0 < a) (hb : 0 < b) : 0 < a ^ b := -ennreal.rpow_pos_of_nonneg ha hb.le - -/-- Auxiliary definition for the `positivity` tactic to handle real powers of extended nonnegative -reals. -/ -meta def prove_ennrpow (a b : expr) : tactic strictness := -do - strictness_a ← core a, - strictness_b ← core b, - match strictness_a, strictness_b with - | positive pa, positive pb := positive <$> mk_app ``ennrpow_pos [pa, pb] - | positive pa, nonnegative pb := positive <$> mk_app ``ennreal.rpow_pos_of_nonneg [pa, pb] - | _, _ := failed -- We already know `0 ≤ x` for all `x : ℝ≥0∞` - end - -end positivity - -open positivity - -/-- Extension for the `positivity` tactic: exponentiation by a real number is nonnegative when the -base is nonnegative and positive when the base is positive. -/ -@[positivity] -meta def positivity_rpow : expr → tactic strictness -| `(@has_pow.pow _ _ real.has_pow %%a %%b) := prove_rpow a b -| `(real.rpow %%a %%b) := prove_rpow a b -| `(@has_pow.pow _ _ nnreal.real.has_pow %%a %%b) := prove_nnrpow a b -| `(nnreal.rpow %%a %%b) := prove_nnrpow a b -| `(@has_pow.pow _ _ ennreal.real.has_pow %%a %%b) := prove_ennrpow a b -| `(ennreal.rpow %%a %%b) := prove_ennrpow a b -| _ := failed - -end tactic diff --git a/src/analysis/special_functions/pow/asymptotics.lean b/src/analysis/special_functions/pow/asymptotics.lean new file mode 100644 index 0000000000000..a906c4f2d5db1 --- /dev/null +++ b/src/analysis/special_functions/pow/asymptotics.lean @@ -0,0 +1,282 @@ +/- +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, Sébastien Gouëzel, + Rémy Degenne, David Loeffler +-/ +import analysis.special_functions.pow.nnreal + +/-! +# Limits and asymptotics of power functions at `+∞` + +This file contains results about the limiting behaviour of power functions at `+∞`. For convenience +some results on asymptotics as `x → 0` (those which are not just continuity statements) are also +located here. +-/ + +noncomputable theory + +open_locale classical real topology nnreal ennreal filter big_operators complex_conjugate +open filter finset set + +/-! +## Limits at `+∞` +-/ + +section limits + +open real filter + +/-- The function `x ^ y` tends to `+∞` at `+∞` for any positive real `y`. -/ +lemma tendsto_rpow_at_top {y : ℝ} (hy : 0 < y) : tendsto (λ x : ℝ, x ^ y) at_top at_top := +begin + rw tendsto_at_top_at_top, + intro b, + use (max b 0) ^ (1/y), + intros x hx, + exact le_of_max_le_left + (by { convert rpow_le_rpow (rpow_nonneg_of_nonneg (le_max_right b 0) (1/y)) hx (le_of_lt hy), + rw [← rpow_mul (le_max_right b 0), (eq_div_iff (ne_of_gt hy)).mp rfl, rpow_one] }), +end + +/-- The function `x ^ (-y)` tends to `0` at `+∞` for any positive real `y`. -/ +lemma tendsto_rpow_neg_at_top {y : ℝ} (hy : 0 < y) : tendsto (λ x : ℝ, x ^ (-y)) at_top (𝓝 0) := +tendsto.congr' (eventually_eq_of_mem (Ioi_mem_at_top 0) (λ x hx, (rpow_neg (le_of_lt hx) y).symm)) + (tendsto_rpow_at_top hy).inv_tendsto_at_top + +/-- The function `x ^ (a / (b * x + c))` tends to `1` at `+∞`, for any real numbers `a`, `b`, and +`c` such that `b` is nonzero. -/ +lemma tendsto_rpow_div_mul_add (a b c : ℝ) (hb : 0 ≠ b) : + tendsto (λ x, x ^ (a / (b*x+c))) at_top (𝓝 1) := +begin + refine tendsto.congr' _ ((tendsto_exp_nhds_0_nhds_1.comp + (by simpa only [mul_zero, pow_one] using ((@tendsto_const_nhds _ _ _ a _).mul + (tendsto_div_pow_mul_exp_add_at_top b c 1 hb)))).comp tendsto_log_at_top), + apply eventually_eq_of_mem (Ioi_mem_at_top (0:ℝ)), + intros x hx, + simp only [set.mem_Ioi, function.comp_app] at hx ⊢, + rw [exp_log hx, ← exp_log (rpow_pos_of_pos hx (a / (b * x + c))), log_rpow hx (a / (b * x + c))], + field_simp, +end + +/-- The function `x ^ (1 / x)` tends to `1` at `+∞`. -/ +lemma tendsto_rpow_div : tendsto (λ x, x ^ ((1:ℝ) / x)) at_top (𝓝 1) := +by { convert tendsto_rpow_div_mul_add (1:ℝ) _ (0:ℝ) zero_ne_one, funext, congr' 2, ring } + +/-- The function `x ^ (-1 / x)` tends to `1` at `+∞`. -/ +lemma tendsto_rpow_neg_div : tendsto (λ x, x ^ (-(1:ℝ) / x)) at_top (𝓝 1) := +by { convert tendsto_rpow_div_mul_add (-(1:ℝ)) _ (0:ℝ) zero_ne_one, funext, congr' 2, ring } + +/-- The function `exp(x) / x ^ s` tends to `+∞` at `+∞`, for any real number `s`. -/ +lemma tendsto_exp_div_rpow_at_top (s : ℝ) : tendsto (λ x : ℝ, exp x / x ^ s) at_top at_top := +begin + cases archimedean_iff_nat_lt.1 (real.archimedean) s with n hn, + refine tendsto_at_top_mono' _ _ (tendsto_exp_div_pow_at_top n), + filter_upwards [eventually_gt_at_top (0 : ℝ), eventually_ge_at_top (1 : ℝ)] with x hx₀ hx₁, + rw [div_le_div_left (exp_pos _) (pow_pos hx₀ _) (rpow_pos_of_pos hx₀ _), ←rpow_nat_cast], + exact rpow_le_rpow_of_exponent_le hx₁ hn.le, +end + +/-- The function `exp (b * x) / x ^ s` tends to `+∞` at `+∞`, for any real `s` and `b > 0`. -/ +lemma tendsto_exp_mul_div_rpow_at_top (s : ℝ) (b : ℝ) (hb : 0 < b) : + tendsto (λ x : ℝ, exp (b * x) / x ^ s) at_top at_top := +begin + refine ((tendsto_rpow_at_top hb).comp (tendsto_exp_div_rpow_at_top (s / b))).congr' _, + filter_upwards [eventually_ge_at_top (0 : ℝ)] with x hx₀, + simp [div_rpow, (exp_pos x).le, rpow_nonneg_of_nonneg, ←rpow_mul, ←exp_mul, mul_comm x, hb.ne', *] +end + +/-- The function `x ^ s * exp (-b * x)` tends to `0` at `+∞`, for any real `s` and `b > 0`. -/ +lemma tendsto_rpow_mul_exp_neg_mul_at_top_nhds_0 (s : ℝ) (b : ℝ) (hb : 0 < b): + tendsto (λ x : ℝ, x ^ s * exp (-b * x)) at_top (𝓝 0) := +begin + refine (tendsto_exp_mul_div_rpow_at_top s b hb).inv_tendsto_at_top.congr' _, + filter_upwards with x using by simp [exp_neg, inv_div, div_eq_mul_inv _ (exp _)] +end + +theorem nnreal.tendsto_rpow_at_top {y : ℝ} (hy : 0 < y) : + tendsto (λ (x : ℝ≥0), x ^ y) at_top at_top := +begin + rw filter.tendsto_at_top_at_top, + intros b, + obtain ⟨c, hc⟩ := tendsto_at_top_at_top.mp (tendsto_rpow_at_top hy) b, + use c.to_nnreal, + intros a ha, + exact_mod_cast hc a (real.to_nnreal_le_iff_le_coe.mp ha), +end + +theorem ennreal.tendsto_rpow_at_top {y : ℝ} (hy : 0 < y) : + tendsto (λ (x : ℝ≥0∞), x ^ y) (𝓝 ⊤) (𝓝 ⊤) := +begin + rw ennreal.tendsto_nhds_top_iff_nnreal, + intros x, + obtain ⟨c, _, hc⟩ := + (at_top_basis_Ioi.tendsto_iff at_top_basis_Ioi).mp (nnreal.tendsto_rpow_at_top hy) x trivial, + have hc' : set.Ioi (↑c) ∈ 𝓝 (⊤ : ℝ≥0∞) := Ioi_mem_nhds ennreal.coe_lt_top, + refine eventually_of_mem hc' _, + intros a ha, + by_cases ha' : a = ⊤, + { simp [ha', hy] }, + lift a to ℝ≥0 using ha', + change ↑c < ↑a at ha, + rw ennreal.coe_rpow_of_nonneg _ hy.le, + exact_mod_cast hc a (by exact_mod_cast ha), +end + +end limits + +/-! +## Asymptotic results: `is_O`, `is_o` and `is_Theta` +-/ +namespace complex +section + +variables {α : Type*} {l : filter α} {f g : α → ℂ} + +open asymptotics + +lemma is_Theta_exp_arg_mul_im (hl : is_bounded_under (≤) l (λ x, |(g x).im|)) : + (λ x, real.exp (arg (f x) * im (g x))) =Θ[l] (λ x, (1 : ℝ)) := +begin + rcases hl with ⟨b, hb⟩, + refine real.is_Theta_exp_comp_one.2 ⟨π * b, _⟩, + rw eventually_map at hb ⊢, + refine hb.mono (λ x hx, _), + erw [abs_mul], + exact mul_le_mul (abs_arg_le_pi _) hx (abs_nonneg _) real.pi_pos.le +end + +lemma is_O_cpow_rpow (hl : is_bounded_under (≤) l (λ x, |(g x).im|)) : + (λ x, f x ^ g x) =O[l] (λ x, abs (f x) ^ (g x).re) := +calc (λ x, f x ^ g x) =O[l] (λ x, abs (f x) ^ (g x).re / real.exp (arg (f x) * im (g x))) : + is_O_of_le _ $ λ x, (abs_cpow_le _ _).trans (le_abs_self _) +... =Θ[l] (λ x, abs (f x) ^ (g x).re / (1 : ℝ)) : + (is_Theta_refl _ _).div (is_Theta_exp_arg_mul_im hl) +... =ᶠ[l] (λ x, abs (f x) ^ (g x).re) : by simp only [of_real_one, div_one] + +lemma is_Theta_cpow_rpow (hl_im : is_bounded_under (≤) l (λ x, |(g x).im|)) + (hl : ∀ᶠ x in l, f x = 0 → re (g x) = 0 → g x = 0): + (λ x, f x ^ g x) =Θ[l] (λ x, abs (f x) ^ (g x).re) := +calc (λ x, f x ^ g x) =Θ[l] (λ x, abs (f x) ^ (g x).re / real.exp (arg (f x) * im (g x))) : + is_Theta_of_norm_eventually_eq' $ hl.mono $ λ x, abs_cpow_of_imp +... =Θ[l] (λ x, abs (f x) ^ (g x).re / (1 : ℝ)) : + (is_Theta_refl _ _).div (is_Theta_exp_arg_mul_im hl_im) +... =ᶠ[l] (λ x, abs (f x) ^ (g x).re) : by simp only [of_real_one, div_one] + +lemma is_Theta_cpow_const_rpow {b : ℂ} (hl : b.re = 0 → b ≠ 0 → ∀ᶠ x in l, f x ≠ 0) : + (λ x, f x ^ b) =Θ[l] (λ x, abs (f x) ^ b.re) := +is_Theta_cpow_rpow is_bounded_under_const $ by simpa only [eventually_imp_distrib_right, ne.def, + ← not_frequently, not_imp_not, imp.swap] using hl + +end + +end complex + +open real + +namespace asymptotics + +variables {α : Type*} {r c : ℝ} {l : filter α} {f g : α → ℝ} + +lemma is_O_with.rpow (h : is_O_with c l f g) (hc : 0 ≤ c) (hr : 0 ≤ r) (hg : 0 ≤ᶠ[l] g) : + is_O_with (c ^ r) l (λ x, f x ^ r) (λ x, g x ^ r) := +begin + apply is_O_with.of_bound, + filter_upwards [hg, h.bound] with x hgx hx, + calc |f x ^ r| ≤ |f x| ^ r : abs_rpow_le_abs_rpow _ _ + ... ≤ (c * |g x|) ^ r : rpow_le_rpow (abs_nonneg _) hx hr + ... = c ^ r * |g x ^ r| : by rw [mul_rpow hc (abs_nonneg _), abs_rpow_of_nonneg hgx] +end + +lemma is_O.rpow (hr : 0 ≤ r) (hg : 0 ≤ᶠ[l] g) (h : f =O[l] g) : + (λ x, f x ^ r) =O[l] (λ x, g x ^ r) := +let ⟨c, hc, h'⟩ := h.exists_nonneg in (h'.rpow hc hr hg).is_O + +lemma is_o.rpow (hr : 0 < r) (hg : 0 ≤ᶠ[l] g) (h : f =o[l] g) : + (λ x, f x ^ r) =o[l] (λ x, g x ^ r) := +is_o.of_is_O_with $ λ c hc, ((h.forall_is_O_with (rpow_pos_of_pos hc r⁻¹)).rpow + (rpow_nonneg_of_nonneg hc.le _) hr.le hg).congr_const + (by rw [←rpow_mul hc.le, inv_mul_cancel hr.ne', rpow_one]) + +end asymptotics + +open asymptotics + +/-- `x ^ s = o(exp(b * x))` as `x → ∞` for any real `s` and positive `b`. -/ +lemma is_o_rpow_exp_pos_mul_at_top (s : ℝ) {b : ℝ} (hb : 0 < b) : + (λ x : ℝ, x ^ s) =o[at_top] (λ x, exp (b * x)) := +iff.mpr (is_o_iff_tendsto $ λ x h, absurd h (exp_pos _).ne') $ + by simpa only [div_eq_mul_inv, exp_neg, neg_mul] + using tendsto_rpow_mul_exp_neg_mul_at_top_nhds_0 s b hb + +/-- `x ^ k = o(exp(b * x))` as `x → ∞` for any integer `k` and positive `b`. -/ +lemma is_o_zpow_exp_pos_mul_at_top (k : ℤ) {b : ℝ} (hb : 0 < b) : + (λ x : ℝ, x ^ k) =o[at_top] (λ x, exp (b * x)) := +by simpa only [rpow_int_cast] using is_o_rpow_exp_pos_mul_at_top k hb + +/-- `x ^ k = o(exp(b * x))` as `x → ∞` for any natural `k` and positive `b`. -/ +lemma is_o_pow_exp_pos_mul_at_top (k : ℕ) {b : ℝ} (hb : 0 < b) : + (λ x : ℝ, x ^ k) =o[at_top] (λ x, exp (b * x)) := +by simpa using is_o_zpow_exp_pos_mul_at_top k hb + +/-- `x ^ s = o(exp x)` as `x → ∞` for any real `s`. -/ +lemma is_o_rpow_exp_at_top (s : ℝ) : (λ x : ℝ, x ^ s) =o[at_top] exp := +by simpa only [one_mul] using is_o_rpow_exp_pos_mul_at_top s one_pos + +/-- `exp (-a * x) = o(x ^ s)` as `x → ∞`, for any positive `a` and real `s`. -/ +lemma is_o_exp_neg_mul_rpow_at_top {a : ℝ} (ha : 0 < a) (b : ℝ) : + is_o at_top (λ x : ℝ, exp (-a * x)) (λ x : ℝ, x ^ b) := +begin + apply is_o_of_tendsto', + { refine (eventually_gt_at_top 0).mp (eventually_of_forall $ λ t ht h, _), + rw rpow_eq_zero_iff_of_nonneg ht.le at h, + exact (ht.ne' h.1).elim }, + { refine (tendsto_exp_mul_div_rpow_at_top (-b) a ha).inv_tendsto_at_top.congr' _, + refine (eventually_ge_at_top 0).mp (eventually_of_forall $ λ t ht, _), + dsimp only, + rw [pi.inv_apply, inv_div, ←inv_div_inv, neg_mul, real.exp_neg, rpow_neg ht, inv_inv] } +end + +lemma is_o_log_rpow_at_top {r : ℝ} (hr : 0 < r) : log =o[at_top] (λ x, x ^ r) := +calc log =O[at_top] (λ x, r * log x) : is_O_self_const_mul _ hr.ne' _ _ + ... =ᶠ[at_top] (λ x, log (x ^ r)) : + (eventually_gt_at_top 0).mono $ λ x hx, (log_rpow hx _).symm + ... =o[at_top] (λ x, x ^ r) : is_o_log_id_at_top.comp_tendsto (tendsto_rpow_at_top hr) + +lemma is_o_log_rpow_rpow_at_top {s : ℝ} (r : ℝ) (hs : 0 < s) : + (λ x, log x ^ r) =o[at_top] (λ x, x ^ s) := +let r' := max r 1 in +have hr : 0 < r', from lt_max_iff.2 $ or.inr one_pos, +have H : 0 < s / r', from div_pos hs hr, +calc (λ x, log x ^ r) =O[at_top] (λ x, log x ^ r') : + is_O.of_bound 1 $ (tendsto_log_at_top.eventually_ge_at_top 1).mono $ λ x hx, + have hx₀ : 0 ≤ log x, from zero_le_one.trans hx, + by simp [norm_eq_abs, abs_rpow_of_nonneg, abs_rpow_of_nonneg hx₀, + rpow_le_rpow_of_exponent_le (hx.trans (le_abs_self _))] + ... =o[at_top] (λ x, (x ^ (s / r')) ^ r') : + (is_o_log_rpow_at_top H).rpow hr $ (tendsto_rpow_at_top H).eventually $ eventually_ge_at_top 0 + ... =ᶠ[at_top] (λ x, x ^ s) : + (eventually_ge_at_top 0).mono $ λ x hx, by simp only [← rpow_mul hx, div_mul_cancel _ hr.ne'] + +lemma is_o_abs_log_rpow_rpow_nhds_zero {s : ℝ} (r : ℝ) (hs : s < 0) : + (λ x, |log x| ^ r) =o[𝓝[>] 0] (λ x, x ^ s) := +((is_o_log_rpow_rpow_at_top r (neg_pos.2 hs)).comp_tendsto tendsto_inv_zero_at_top).congr' + (mem_of_superset (Icc_mem_nhds_within_Ioi $ set.left_mem_Ico.2 one_pos) $ + λ x hx, by simp [abs_of_nonpos, log_nonpos hx.1 hx.2]) + (eventually_mem_nhds_within.mono $ λ x hx, + by rw [function.comp_app, inv_rpow hx.out.le, rpow_neg hx.out.le, inv_inv]) + +lemma is_o_log_rpow_nhds_zero {r : ℝ} (hr : r < 0) : log =o[𝓝[>] 0] (λ x, x ^ r) := +(is_o_abs_log_rpow_rpow_nhds_zero 1 hr).neg_left.congr' + (mem_of_superset (Icc_mem_nhds_within_Ioi $ set.left_mem_Ico.2 one_pos) $ + λ x hx, by simp [abs_of_nonpos (log_nonpos hx.1 hx.2)]) + eventually_eq.rfl + +lemma tendsto_log_div_rpow_nhds_zero {r : ℝ} (hr : r < 0) : + tendsto (λ x, log x / x ^ r) (𝓝[>] 0) (𝓝 0) := +(is_o_log_rpow_nhds_zero hr).tendsto_div_nhds_zero + +lemma tendsto_log_mul_rpow_nhds_zero {r : ℝ} (hr : 0 < r) : + tendsto (λ x, log x * x ^ r) (𝓝[>] 0) (𝓝 0) := +(tendsto_log_div_rpow_nhds_zero $ neg_lt_zero.2 hr).congr' $ + eventually_mem_nhds_within.mono $ λ x hx, by rw [rpow_neg hx.out.le, div_inv_eq_mul] diff --git a/src/analysis/special_functions/pow/complex.lean b/src/analysis/special_functions/pow/complex.lean new file mode 100644 index 0000000000000..d60d119b4a50b --- /dev/null +++ b/src/analysis/special_functions/pow/complex.lean @@ -0,0 +1,179 @@ +/- +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, Sébastien Gouëzel, + Rémy Degenne, David Loeffler +-/ +import analysis.special_functions.complex.log + +/-! # Power function on `ℂ` + +We construct the power functions `x ^ y`, where `x` and `y` are complex numbers. +-/ + +open_locale classical real topology filter complex_conjugate +open filter finset set + +namespace complex + +/-- The complex power function `x ^ y`, given by `x ^ y = exp(y log x)` (where `log` is the +principal determination of the logarithm), unless `x = 0` where one sets `0 ^ 0 = 1` and +`0 ^ y = 0` for `y ≠ 0`. -/ +noncomputable def cpow (x y : ℂ) : ℂ := +if x = 0 + then if y = 0 + then 1 + else 0 + else exp (log x * y) + +noncomputable instance : has_pow ℂ ℂ := ⟨cpow⟩ + +@[simp] lemma cpow_eq_pow (x y : ℂ) : cpow x y = x ^ y := rfl + +lemma cpow_def (x y : ℂ) : x ^ y = + if x = 0 + then if y = 0 + then 1 + else 0 + else exp (log x * y) := rfl + +lemma cpow_def_of_ne_zero {x : ℂ} (hx : x ≠ 0) (y : ℂ) : x ^ y = exp (log x * y) := if_neg hx + +@[simp] lemma cpow_zero (x : ℂ) : x ^ (0 : ℂ) = 1 := by simp [cpow_def] + +@[simp] lemma cpow_eq_zero_iff (x y : ℂ) : x ^ y = 0 ↔ x = 0 ∧ y ≠ 0 := +by { simp only [cpow_def], split_ifs; simp [*, exp_ne_zero] } + +@[simp] lemma zero_cpow {x : ℂ} (h : x ≠ 0) : (0 : ℂ) ^ x = 0 := +by simp [cpow_def, *] + +lemma zero_cpow_eq_iff {x : ℂ} {a : ℂ} : 0 ^ x = a ↔ (x ≠ 0 ∧ a = 0) ∨ (x = 0 ∧ a = 1) := +begin + split, + { intros hyp, + simp only [cpow_def, eq_self_iff_true, if_true] at hyp, + by_cases x = 0, + { subst h, simp only [if_true, eq_self_iff_true] at hyp, right, exact ⟨rfl, hyp.symm⟩}, + { rw if_neg h at hyp, left, exact ⟨h, hyp.symm⟩, }, }, + { rintro (⟨h, rfl⟩|⟨rfl,rfl⟩), + { exact zero_cpow h, }, + { exact cpow_zero _, }, }, +end + +lemma eq_zero_cpow_iff {x : ℂ} {a : ℂ} : a = 0 ^ x ↔ (x ≠ 0 ∧ a = 0) ∨ (x = 0 ∧ a = 1) := +by rw [←zero_cpow_eq_iff, eq_comm] + +@[simp] lemma cpow_one (x : ℂ) : x ^ (1 : ℂ) = x := +if hx : x = 0 then by simp [hx, cpow_def] +else by rw [cpow_def, if_neg (one_ne_zero : (1 : ℂ) ≠ 0), if_neg hx, mul_one, exp_log hx] + +@[simp] lemma one_cpow (x : ℂ) : (1 : ℂ) ^ x = 1 := +by rw cpow_def; split_ifs; simp [one_ne_zero, *] at * + +lemma cpow_add {x : ℂ} (y z : ℂ) (hx : x ≠ 0) : x ^ (y + z) = x ^ y * x ^ z := +by simp only [cpow_def, ite_mul, boole_mul, mul_ite, mul_boole]; simp [*, exp_add, mul_add] at * + +lemma cpow_mul {x y : ℂ} (z : ℂ) (h₁ : -π < (log x * y).im) (h₂ : (log x * y).im ≤ π) : + x ^ (y * z) = (x ^ y) ^ z := +begin + simp only [cpow_def], + split_ifs; + simp [*, exp_ne_zero, log_exp h₁ h₂, mul_assoc] at * +end + +lemma cpow_neg (x y : ℂ) : x ^ -y = (x ^ y)⁻¹ := +by simp only [cpow_def, neg_eq_zero, mul_neg]; split_ifs; simp [exp_neg] + +lemma cpow_sub {x : ℂ} (y z : ℂ) (hx : x ≠ 0) : x ^ (y - z) = x ^ y / x ^ z := +by rw [sub_eq_add_neg, cpow_add _ _ hx, cpow_neg, div_eq_mul_inv] + +lemma cpow_neg_one (x : ℂ) : x ^ (-1 : ℂ) = x⁻¹ := +by simpa using cpow_neg x 1 + +@[simp, norm_cast] lemma cpow_nat_cast (x : ℂ) : ∀ (n : ℕ), x ^ (n : ℂ) = x ^ n +| 0 := by simp +| (n + 1) := if hx : x = 0 then by simp only [hx, pow_succ, + complex.zero_cpow (nat.cast_ne_zero.2 (nat.succ_ne_zero _)), zero_mul] + else by simp [cpow_add, hx, pow_add, cpow_nat_cast n] + +@[simp] lemma cpow_two (x : ℂ) : x ^ (2 : ℂ) = x ^ 2 := +by { rw ← cpow_nat_cast, simp only [nat.cast_bit0, nat.cast_one] } + +@[simp, norm_cast] lemma cpow_int_cast (x : ℂ) : ∀ (n : ℤ), x ^ (n : ℂ) = x ^ n +| (n : ℕ) := by simp +| -[1+ n] := by rw zpow_neg_succ_of_nat; + simp only [int.neg_succ_of_nat_coe, int.cast_neg, complex.cpow_neg, inv_eq_one_div, + int.cast_coe_nat, cpow_nat_cast] + +lemma cpow_nat_inv_pow (x : ℂ) {n : ℕ} (hn : n ≠ 0) : (x ^ (n⁻¹ : ℂ)) ^ n = x := +begin + suffices : im (log x * n⁻¹) ∈ Ioc (-π) π, + { rw [← cpow_nat_cast, ← cpow_mul _ this.1 this.2, inv_mul_cancel, cpow_one], + exact_mod_cast hn }, + rw [mul_comm, ← of_real_nat_cast, ← of_real_inv, of_real_mul_im, ← div_eq_inv_mul], + rw [← pos_iff_ne_zero] at hn, + have hn' : 0 < (n : ℝ), by assumption_mod_cast, + have hn1 : 1 ≤ (n : ℝ), by exact_mod_cast (nat.succ_le_iff.2 hn), + split, + { rw lt_div_iff hn', + calc -π * n ≤ -π * 1 : mul_le_mul_of_nonpos_left hn1 (neg_nonpos.2 real.pi_pos.le) + ... = -π : mul_one _ + ... < im (log x) : neg_pi_lt_log_im _ }, + { rw div_le_iff hn', + calc im (log x) ≤ π : log_im_le_pi _ + ... = π * 1 : (mul_one π).symm + ... ≤ π * n : mul_le_mul_of_nonneg_left hn1 real.pi_pos.le } +end + +lemma mul_cpow_of_real_nonneg {a b : ℝ} (ha : 0 ≤ a) (hb : 0 ≤ b) (r : ℂ) : + ((a : ℂ) * (b : ℂ)) ^ r = (a : ℂ) ^ r * (b : ℂ) ^ r := +begin + rcases eq_or_ne r 0 with rfl | hr, + { simp only [cpow_zero, mul_one] }, + rcases eq_or_lt_of_le ha with rfl | ha', + { rw [of_real_zero, zero_mul, zero_cpow hr, zero_mul] }, + rcases eq_or_lt_of_le hb with rfl | hb', + { rw [of_real_zero, mul_zero, zero_cpow hr, mul_zero] }, + have ha'' : (a : ℂ) ≠ 0 := of_real_ne_zero.mpr ha'.ne', + have hb'' : (b : ℂ) ≠ 0 := of_real_ne_zero.mpr hb'.ne', + rw [cpow_def_of_ne_zero (mul_ne_zero ha'' hb''), log_of_real_mul ha' hb'', of_real_log ha, + add_mul, exp_add, ←cpow_def_of_ne_zero ha'', ←cpow_def_of_ne_zero hb''] +end + +lemma inv_cpow_eq_ite (x : ℂ) (n : ℂ) : + x⁻¹ ^ n = if x.arg = π then conj (x ^ conj n)⁻¹ else (x ^ n)⁻¹ := +begin + simp_rw [complex.cpow_def, log_inv_eq_ite, inv_eq_zero, map_eq_zero, ite_mul, neg_mul, + is_R_or_C.conj_inv, apply_ite conj, apply_ite exp, apply_ite has_inv.inv, map_zero, map_one, + exp_neg, inv_one, inv_zero, ←exp_conj, map_mul, conj_conj], + split_ifs with hx hn ha ha; refl, +end + +lemma inv_cpow (x : ℂ) (n : ℂ) (hx : x.arg ≠ π) : x⁻¹ ^ n = (x ^ n)⁻¹ := +by rw [inv_cpow_eq_ite, if_neg hx] + +/-- `complex.inv_cpow_eq_ite` with the `ite` on the other side. -/ +lemma inv_cpow_eq_ite' (x : ℂ) (n : ℂ) : + (x ^ n)⁻¹ = if x.arg = π then conj (x⁻¹ ^ conj n) else x⁻¹ ^ n := +begin + rw [inv_cpow_eq_ite, apply_ite conj, conj_conj, conj_conj], + split_ifs, + { refl }, + { rw inv_cpow _ _ h } +end + +lemma conj_cpow_eq_ite (x : ℂ) (n : ℂ) : + conj x ^ n = if x.arg = π then x ^ n else conj (x ^ conj n) := +begin + simp_rw [cpow_def, map_eq_zero, apply_ite conj, map_one, map_zero, ←exp_conj, map_mul, + conj_conj, log_conj_eq_ite], + split_ifs with hcx hn hx; refl +end + +lemma conj_cpow (x : ℂ) (n : ℂ) (hx : x.arg ≠ π) : conj x ^ n = conj (x ^ conj n) := +by rw [conj_cpow_eq_ite, if_neg hx] + +lemma cpow_conj (x : ℂ) (n : ℂ) (hx : x.arg ≠ π) : x ^ conj n = conj (conj x ^ n) := +by rw [conj_cpow _ _ hx, conj_conj] + +end complex diff --git a/src/analysis/special_functions/pow/continuity.lean b/src/analysis/special_functions/pow/continuity.lean new file mode 100644 index 0000000000000..a970291facd82 --- /dev/null +++ b/src/analysis/special_functions/pow/continuity.lean @@ -0,0 +1,500 @@ +/- +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, Sébastien Gouëzel, + Rémy Degenne, David Loeffler +-/ +import analysis.special_functions.pow.asymptotics + +/-! +# Continuity of power functions + +This file contains lemmas about continuity of the power functions on `ℂ`, `ℝ`, `ℝ≥0`, and `ℝ≥0∞`. +-/ + +noncomputable theory + +open_locale classical real topology nnreal ennreal filter big_operators complex_conjugate +open filter finset set + +section cpow_limits + +/-! +## Continuity for complex powers +-/ + +open complex + +variables {α : Type*} + +lemma zero_cpow_eq_nhds {b : ℂ} (hb : b ≠ 0) : + (λ (x : ℂ), (0 : ℂ) ^ x) =ᶠ[𝓝 b] 0 := +begin + suffices : ∀ᶠ (x : ℂ) in (𝓝 b), x ≠ 0, + from this.mono (λ x hx, by { dsimp only, rw [zero_cpow hx, pi.zero_apply]} ), + exact is_open.eventually_mem is_open_ne hb, +end + +lemma cpow_eq_nhds {a b : ℂ} (ha : a ≠ 0) : + (λ x, x ^ b) =ᶠ[𝓝 a] λ x, exp (log x * b) := +begin + suffices : ∀ᶠ (x : ℂ) in (𝓝 a), x ≠ 0, + from this.mono (λ x hx, by { dsimp only, rw [cpow_def_of_ne_zero hx], }), + exact is_open.eventually_mem is_open_ne ha, +end + +lemma cpow_eq_nhds' {p : ℂ × ℂ} (hp_fst : p.fst ≠ 0) : + (λ x, x.1 ^ x.2) =ᶠ[𝓝 p] λ x, exp (log x.1 * x.2) := +begin + suffices : ∀ᶠ (x : ℂ × ℂ) in (𝓝 p), x.1 ≠ 0, + from this.mono (λ x hx, by { dsimp only, rw cpow_def_of_ne_zero hx, }), + refine is_open.eventually_mem _ hp_fst, + change is_open {x : ℂ × ℂ | x.1 = 0}ᶜ, + rw is_open_compl_iff, + exact is_closed_eq continuous_fst continuous_const, +end + +/- Continuity of `λ x, a ^ x`: union of these two lemmas is optimal. -/ + +lemma continuous_at_const_cpow {a b : ℂ} (ha : a ≠ 0) : continuous_at (λ x, a ^ x) b := +begin + have cpow_eq : (λ x:ℂ, a ^ x) = λ x, exp (log a * x), + by { ext1 b, rw [cpow_def_of_ne_zero ha], }, + rw cpow_eq, + exact continuous_exp.continuous_at.comp (continuous_at.mul continuous_at_const continuous_at_id), +end + +lemma continuous_at_const_cpow' {a b : ℂ} (h : b ≠ 0) : continuous_at (λ x, a ^ x) b := +begin + by_cases ha : a = 0, + { rw [ha, continuous_at_congr (zero_cpow_eq_nhds h)], exact continuous_at_const, }, + { exact continuous_at_const_cpow ha, }, +end + +/-- The function `z ^ w` is continuous in `(z, w)` provided that `z` does not belong to the interval +`(-∞, 0]` on the real line. See also `complex.continuous_at_cpow_zero_of_re_pos` for a version that +works for `z = 0` but assumes `0 < re w`. -/ +lemma continuous_at_cpow {p : ℂ × ℂ} (hp_fst : 0 < p.fst.re ∨ p.fst.im ≠ 0) : + continuous_at (λ x : ℂ × ℂ, x.1 ^ x.2) p := +begin + have hp_fst_ne_zero : p.fst ≠ 0, + by { intro h, cases hp_fst; { rw h at hp_fst, simpa using hp_fst, }, }, + rw continuous_at_congr (cpow_eq_nhds' hp_fst_ne_zero), + refine continuous_exp.continuous_at.comp _, + refine continuous_at.mul (continuous_at.comp _ continuous_fst.continuous_at) + continuous_snd.continuous_at, + exact continuous_at_clog hp_fst, +end + +lemma continuous_at_cpow_const {a b : ℂ} (ha : 0 < a.re ∨ a.im ≠ 0) : + continuous_at (λ x, cpow x b) a := +tendsto.comp (@continuous_at_cpow (a, b) ha) (continuous_at_id.prod continuous_at_const) + +lemma filter.tendsto.cpow {l : filter α} {f g : α → ℂ} {a b : ℂ} (hf : tendsto f l (𝓝 a)) + (hg : tendsto g l (𝓝 b)) (ha : 0 < a.re ∨ a.im ≠ 0) : + tendsto (λ x, f x ^ g x) l (𝓝 (a ^ b)) := +(@continuous_at_cpow (a,b) ha).tendsto.comp (hf.prod_mk_nhds hg) + +lemma filter.tendsto.const_cpow {l : filter α} {f : α → ℂ} {a b : ℂ} (hf : tendsto f l (𝓝 b)) + (h : a ≠ 0 ∨ b ≠ 0) : + tendsto (λ x, a ^ f x) l (𝓝 (a ^ b)) := +begin + cases h, + { exact (continuous_at_const_cpow h).tendsto.comp hf, }, + { exact (continuous_at_const_cpow' h).tendsto.comp hf, }, +end + +variables [topological_space α] {f g : α → ℂ} {s : set α} {a : α} + +lemma continuous_within_at.cpow (hf : continuous_within_at f s a) (hg : continuous_within_at g s a) + (h0 : 0 < (f a).re ∨ (f a).im ≠ 0) : + continuous_within_at (λ x, f x ^ g x) s a := +hf.cpow hg h0 + +lemma continuous_within_at.const_cpow {b : ℂ} (hf : continuous_within_at f s a) + (h : b ≠ 0 ∨ f a ≠ 0) : + continuous_within_at (λ x, b ^ f x) s a := +hf.const_cpow h + +lemma continuous_at.cpow (hf : continuous_at f a) (hg : continuous_at g a) + (h0 : 0 < (f a).re ∨ (f a).im ≠ 0) : + continuous_at (λ x, f x ^ g x) a := +hf.cpow hg h0 + +lemma continuous_at.const_cpow {b : ℂ} (hf : continuous_at f a) (h : b ≠ 0 ∨ f a ≠ 0) : + continuous_at (λ x, b ^ f x) a := +hf.const_cpow h + +lemma continuous_on.cpow (hf : continuous_on f s) (hg : continuous_on g s) + (h0 : ∀ a ∈ s, 0 < (f a).re ∨ (f a).im ≠ 0) : + continuous_on (λ x, f x ^ g x) s := +λ a ha, (hf a ha).cpow (hg a ha) (h0 a ha) + +lemma continuous_on.const_cpow {b : ℂ} (hf : continuous_on f s) (h : b ≠ 0 ∨ ∀ a ∈ s, f a ≠ 0) : + continuous_on (λ x, b ^ f x) s := +λ a ha, (hf a ha).const_cpow (h.imp id $ λ h, h a ha) + +lemma continuous.cpow (hf : continuous f) (hg : continuous g) + (h0 : ∀ a, 0 < (f a).re ∨ (f a).im ≠ 0) : + continuous (λ x, f x ^ g x) := +continuous_iff_continuous_at.2 $ λ a, (hf.continuous_at.cpow hg.continuous_at (h0 a)) + +lemma continuous.const_cpow {b : ℂ} (hf : continuous f) (h : b ≠ 0 ∨ ∀ a, f a ≠ 0) : + continuous (λ x, b ^ f x) := +continuous_iff_continuous_at.2 $ λ a, (hf.continuous_at.const_cpow $ h.imp id $ λ h, h a) + +lemma continuous_on.cpow_const {b : ℂ} (hf : continuous_on f s) + (h : ∀ (a : α), a ∈ s → 0 < (f a).re ∨ (f a).im ≠ 0) : + continuous_on (λ x, (f x) ^ b) s := +hf.cpow continuous_on_const h + +end cpow_limits + +section rpow_limits + +/-! +## Continuity for real powers +-/ + +namespace real + +lemma continuous_at_const_rpow {a b : ℝ} (h : a ≠ 0) : continuous_at (rpow a) b := +begin + have : rpow a = λ x : ℝ, ((a : ℂ) ^ (x : ℂ)).re, by { ext1 x, rw [rpow_eq_pow, rpow_def], }, + rw this, + refine complex.continuous_re.continuous_at.comp _, + refine (continuous_at_const_cpow _).comp complex.continuous_of_real.continuous_at, + norm_cast, + exact h, +end + +lemma continuous_at_const_rpow' {a b : ℝ} (h : b ≠ 0) : continuous_at (rpow a) b := +begin + have : rpow a = λ x : ℝ, ((a : ℂ) ^ (x : ℂ)).re, by { ext1 x, rw [rpow_eq_pow, rpow_def], }, + rw this, + refine complex.continuous_re.continuous_at.comp _, + refine (continuous_at_const_cpow' _).comp complex.continuous_of_real.continuous_at, + norm_cast, + exact h, +end + +lemma rpow_eq_nhds_of_neg {p : ℝ × ℝ} (hp_fst : p.fst < 0) : + (λ x : ℝ × ℝ, x.1 ^ x.2) =ᶠ[𝓝 p] λ x, exp (log x.1 * x.2) * cos (x.2 * π) := +begin + suffices : ∀ᶠ (x : ℝ × ℝ) in (𝓝 p), x.1 < 0, + from this.mono (λ x hx, by { dsimp only, rw rpow_def_of_neg hx, }), + exact is_open.eventually_mem (is_open_lt continuous_fst continuous_const) hp_fst, +end + +lemma rpow_eq_nhds_of_pos {p : ℝ × ℝ} (hp_fst : 0 < p.fst) : + (λ x : ℝ × ℝ, x.1 ^ x.2) =ᶠ[𝓝 p] λ x, exp (log x.1 * x.2) := +begin + suffices : ∀ᶠ (x : ℝ × ℝ) in (𝓝 p), 0 < x.1, + from this.mono (λ x hx, by { dsimp only, rw rpow_def_of_pos hx, }), + exact is_open.eventually_mem (is_open_lt continuous_const continuous_fst) hp_fst, +end + +lemma continuous_at_rpow_of_ne (p : ℝ × ℝ) (hp : p.1 ≠ 0) : + continuous_at (λ p : ℝ × ℝ, p.1 ^ p.2) p := +begin + rw ne_iff_lt_or_gt at hp, + cases hp, + { rw continuous_at_congr (rpow_eq_nhds_of_neg hp), + refine continuous_at.mul _ (continuous_cos.continuous_at.comp _), + { refine continuous_exp.continuous_at.comp (continuous_at.mul _ continuous_snd.continuous_at), + refine (continuous_at_log _).comp continuous_fst.continuous_at, + exact hp.ne, }, + { exact continuous_snd.continuous_at.mul continuous_at_const, }, }, + { rw continuous_at_congr (rpow_eq_nhds_of_pos hp), + refine continuous_exp.continuous_at.comp (continuous_at.mul _ continuous_snd.continuous_at), + refine (continuous_at_log _).comp continuous_fst.continuous_at, + exact hp.lt.ne.symm, }, +end + +lemma continuous_at_rpow_of_pos (p : ℝ × ℝ) (hp : 0 < p.2) : + continuous_at (λ p : ℝ × ℝ, p.1 ^ p.2) p := +begin + cases p with x y, + obtain hx|rfl := ne_or_eq x 0, + { exact continuous_at_rpow_of_ne (x, y) hx }, + have A : tendsto (λ p : ℝ × ℝ, exp (log p.1 * p.2)) (𝓝[≠] 0 ×ᶠ 𝓝 y) (𝓝 0) := + tendsto_exp_at_bot.comp + ((tendsto_log_nhds_within_zero.comp tendsto_fst).at_bot_mul hp tendsto_snd), + have B : tendsto (λ p : ℝ × ℝ, p.1 ^ p.2) (𝓝[≠] 0 ×ᶠ 𝓝 y) (𝓝 0) := + squeeze_zero_norm (λ p, abs_rpow_le_exp_log_mul p.1 p.2) A, + have C : tendsto (λ p : ℝ × ℝ, p.1 ^ p.2) (𝓝[{0}] 0 ×ᶠ 𝓝 y) (pure 0), + { rw [nhds_within_singleton, tendsto_pure, pure_prod, eventually_map], + exact (lt_mem_nhds hp).mono (λ y hy, zero_rpow hy.ne') }, + simpa only [← sup_prod, ← nhds_within_union, compl_union_self, nhds_within_univ, nhds_prod_eq, + continuous_at, zero_rpow hp.ne'] using B.sup (C.mono_right (pure_le_nhds _)) +end + +lemma continuous_at_rpow (p : ℝ × ℝ) (h : p.1 ≠ 0 ∨ 0 < p.2) : + continuous_at (λ p : ℝ × ℝ, p.1 ^ p.2) p := +h.elim (λ h, continuous_at_rpow_of_ne p h) (λ h, continuous_at_rpow_of_pos p h) + +lemma continuous_at_rpow_const (x : ℝ) (q : ℝ) (h : x ≠ 0 ∨ 0 < q) : + continuous_at (λ (x : ℝ), x ^ q) x := +begin + change continuous_at ((λ p : ℝ × ℝ, p.1 ^ p.2) ∘ (λ y : ℝ, (y, q))) x, + apply continuous_at.comp, + { exact continuous_at_rpow (x, q) h }, + { exact (continuous_id'.prod_mk continuous_const).continuous_at } +end + +end real + +section + +variable {α : Type*} + +lemma filter.tendsto.rpow {l : filter α} {f g : α → ℝ} {x y : ℝ} + (hf : tendsto f l (𝓝 x)) (hg : tendsto g l (𝓝 y)) (h : x ≠ 0 ∨ 0 < y) : + tendsto (λ t, f t ^ g t) l (𝓝 (x ^ y)) := +(real.continuous_at_rpow (x, y) h).tendsto.comp (hf.prod_mk_nhds hg) + +lemma filter.tendsto.rpow_const {l : filter α} {f : α → ℝ} {x p : ℝ} + (hf : tendsto f l (𝓝 x)) (h : x ≠ 0 ∨ 0 ≤ p) : + tendsto (λ a, f a ^ p) l (𝓝 (x ^ p)) := +if h0 : 0 = p then h0 ▸ by simp [tendsto_const_nhds] +else hf.rpow tendsto_const_nhds (h.imp id $ λ h', h'.lt_of_ne h0) + +variables [topological_space α] {f g : α → ℝ} {s : set α} {x : α} {p : ℝ} + +lemma continuous_at.rpow (hf : continuous_at f x) (hg : continuous_at g x) (h : f x ≠ 0 ∨ 0 < g x) : + continuous_at (λ t, f t ^ g t) x := +hf.rpow hg h + +lemma continuous_within_at.rpow (hf : continuous_within_at f s x) (hg : continuous_within_at g s x) + (h : f x ≠ 0 ∨ 0 < g x) : + continuous_within_at (λ t, f t ^ g t) s x := +hf.rpow hg h + +lemma continuous_on.rpow (hf : continuous_on f s) (hg : continuous_on g s) + (h : ∀ x ∈ s, f x ≠ 0 ∨ 0 < g x) : + continuous_on (λ t, f t ^ g t) s := +λ t ht, (hf t ht).rpow (hg t ht) (h t ht) + +lemma continuous.rpow (hf : continuous f) (hg : continuous g) (h : ∀ x, f x ≠ 0 ∨ 0 < g x) : + continuous (λ x, f x ^ g x) := +continuous_iff_continuous_at.2 $ λ x, (hf.continuous_at.rpow hg.continuous_at (h x)) + +lemma continuous_within_at.rpow_const (hf : continuous_within_at f s x) (h : f x ≠ 0 ∨ 0 ≤ p) : + continuous_within_at (λ x, f x ^ p) s x := +hf.rpow_const h + +lemma continuous_at.rpow_const (hf : continuous_at f x) (h : f x ≠ 0 ∨ 0 ≤ p) : + continuous_at (λ x, f x ^ p) x := +hf.rpow_const h + +lemma continuous_on.rpow_const (hf : continuous_on f s) (h : ∀ x ∈ s, f x ≠ 0 ∨ 0 ≤ p) : + continuous_on (λ x, f x ^ p) s := +λ x hx, (hf x hx).rpow_const (h x hx) + +lemma continuous.rpow_const (hf : continuous f) (h : ∀ x, f x ≠ 0 ∨ 0 ≤ p) : + continuous (λ x, f x ^ p) := +continuous_iff_continuous_at.2 $ λ x, hf.continuous_at.rpow_const (h x) + +end + +end rpow_limits + +/-! ## Continuity results for `cpow`, part II + +These results involve relating real and complex powers, so cannot be done higher up. +-/ +section cpow_limits2 + +namespace complex + +/-- See also `continuous_at_cpow` and `complex.continuous_at_cpow_of_re_pos`. -/ +lemma continuous_at_cpow_zero_of_re_pos {z : ℂ} (hz : 0 < z.re) : + continuous_at (λ x : ℂ × ℂ, x.1 ^ x.2) (0, z) := +begin + have hz₀ : z ≠ 0, from ne_of_apply_ne re hz.ne', + rw [continuous_at, zero_cpow hz₀, tendsto_zero_iff_norm_tendsto_zero], + refine squeeze_zero (λ _, norm_nonneg _) (λ _, abs_cpow_le _ _) _, + simp only [div_eq_mul_inv, ← real.exp_neg], + refine tendsto.zero_mul_is_bounded_under_le _ _, + { convert (continuous_fst.norm.tendsto _).rpow ((continuous_re.comp continuous_snd).tendsto _) _; + simp [hz, real.zero_rpow hz.ne'] }, + { simp only [(∘), real.norm_eq_abs, abs_of_pos (real.exp_pos _)], + rcases exists_gt (|im z|) with ⟨C, hC⟩, + refine ⟨real.exp (π * C), eventually_map.2 _⟩, + refine (((continuous_im.comp continuous_snd).abs.tendsto (_, z)).eventually + (gt_mem_nhds hC)).mono (λ z hz, real.exp_le_exp.2 $ (neg_le_abs_self _).trans _), + rw _root_.abs_mul, + exact mul_le_mul (abs_le.2 ⟨(neg_pi_lt_arg _).le, arg_le_pi _⟩) hz.le + (_root_.abs_nonneg _) real.pi_pos.le } +end + +/-- See also `continuous_at_cpow` for a version that assumes `p.1 ≠ 0` but makes no +assumptions about `p.2`. -/ +lemma continuous_at_cpow_of_re_pos {p : ℂ × ℂ} (h₁ : 0 ≤ p.1.re ∨ p.1.im ≠ 0) (h₂ : 0 < p.2.re) : + continuous_at (λ x : ℂ × ℂ, x.1 ^ x.2) p := +begin + cases p with z w, + rw [← not_lt_zero_iff, lt_iff_le_and_ne, not_and_distrib, ne.def, not_not, not_le_zero_iff] at h₁, + rcases h₁ with h₁|(rfl : z = 0), + exacts [continuous_at_cpow h₁, continuous_at_cpow_zero_of_re_pos h₂] +end + +/-- See also `continuous_at_cpow_const` for a version that assumes `z ≠ 0` but makes no +assumptions about `w`. -/ +lemma continuous_at_cpow_const_of_re_pos {z w : ℂ} (hz : 0 ≤ re z ∨ im z ≠ 0) (hw : 0 < re w) : + continuous_at (λ x, x ^ w) z := +tendsto.comp (@continuous_at_cpow_of_re_pos (z, w) hz hw) + (continuous_at_id.prod continuous_at_const) + +/-- Continuity of `(x, y) ↦ x ^ y` as a function on `ℝ × ℂ`. -/ +lemma continuous_at_of_real_cpow (x : ℝ) (y : ℂ) (h : 0 < y.re ∨ x ≠ 0) : + continuous_at (λ p, ↑p.1 ^ p.2 : ℝ × ℂ → ℂ) (x, y) := +begin + rcases lt_trichotomy 0 x with hx | rfl | hx, + { -- x > 0 : easy case + have : continuous_at (λ p, ⟨↑p.1, p.2⟩ : ℝ × ℂ → ℂ × ℂ) (x, y), + from continuous_of_real.continuous_at.prod_map continuous_at_id, + refine (continuous_at_cpow (or.inl _)).comp this, + rwa of_real_re }, + { -- x = 0 : reduce to continuous_at_cpow_zero_of_re_pos + have A : continuous_at (λ p, p.1 ^ p.2 : ℂ × ℂ → ℂ) ⟨↑(0:ℝ), y⟩, + { rw of_real_zero, + apply continuous_at_cpow_zero_of_re_pos, + tauto }, + have B : continuous_at (λ p, ⟨↑p.1, p.2⟩ : ℝ × ℂ → ℂ × ℂ) ⟨0, y⟩, + from continuous_of_real.continuous_at.prod_map continuous_at_id, + exact @continuous_at.comp (ℝ × ℂ) (ℂ × ℂ) ℂ _ _ _ _ (λ p, ⟨↑p.1, p.2⟩) ⟨0, y⟩ A B }, + { -- x < 0 : difficult case + suffices : continuous_at (λ p, (-↑p.1) ^ p.2 * exp (π * I * p.2) : ℝ × ℂ → ℂ) (x, y), + { refine this.congr (eventually_of_mem (prod_mem_nhds (Iio_mem_nhds hx) univ_mem) _), + exact λ p hp, (of_real_cpow_of_nonpos (le_of_lt hp.1) p.2).symm }, + have A : continuous_at (λ p, ⟨-↑p.1, p.2⟩ : ℝ × ℂ → ℂ × ℂ) (x, y), + from continuous_at.prod_map (continuous_of_real.continuous_at.neg) continuous_at_id, + apply continuous_at.mul, + { refine (continuous_at_cpow (or.inl _)).comp A, + rwa [neg_re, of_real_re, neg_pos] }, + { exact (continuous_exp.comp (continuous_const.mul continuous_snd)).continuous_at } }, +end + +lemma continuous_at_of_real_cpow_const (x : ℝ) (y : ℂ) (h : 0 < y.re ∨ x ≠ 0) : + continuous_at (λ a, a ^ y : ℝ → ℂ) x := +@continuous_at.comp _ _ _ _ _ _ _ _ x (continuous_at_of_real_cpow x y h) + (continuous_id.prod_mk continuous_const).continuous_at + +lemma continuous_of_real_cpow_const {y : ℂ} (hs : 0 < y.re) : continuous (λ x, x ^ y : ℝ → ℂ) := +continuous_iff_continuous_at.mpr (λ x, continuous_at_of_real_cpow_const x y (or.inl hs)) + +end complex + +end cpow_limits2 + +/-! ## Limits and continuity for `ℝ≥0` powers -/ +namespace nnreal + +lemma continuous_at_rpow {x : ℝ≥0} {y : ℝ} (h : x ≠ 0 ∨ 0 < y) : + continuous_at (λp:ℝ≥0×ℝ, p.1^p.2) (x, y) := +begin + have : (λp:ℝ≥0×ℝ, p.1^p.2) = real.to_nnreal ∘ (λp:ℝ×ℝ, p.1^p.2) ∘ (λp:ℝ≥0 × ℝ, (p.1.1, p.2)), + { ext p, + rw [coe_rpow, real.coe_to_nnreal _ (real.rpow_nonneg_of_nonneg p.1.2 _)], + refl }, + rw this, + refine continuous_real_to_nnreal.continuous_at.comp (continuous_at.comp _ _), + { apply real.continuous_at_rpow, + simp only [ne.def] at h, + rw ← (nnreal.coe_eq_zero x) at h, + exact h }, + { exact ((continuous_subtype_val.comp continuous_fst).prod_mk continuous_snd).continuous_at } +end + +lemma eventually_pow_one_div_le (x : ℝ≥0) {y : ℝ≥0} (hy : 1 < y) : + ∀ᶠ (n : ℕ) in at_top, x ^ (1 / n : ℝ) ≤ y := +begin + obtain ⟨m, hm⟩ := add_one_pow_unbounded_of_pos x (tsub_pos_of_lt hy), + rw [tsub_add_cancel_of_le hy.le] at hm, + refine eventually_at_top.2 ⟨m + 1, λ n hn, _⟩, + simpa only [nnreal.rpow_one_div_le_iff (nat.cast_pos.2 $ m.succ_pos.trans_le hn), + nnreal.rpow_nat_cast] using hm.le.trans (pow_le_pow hy.le (m.le_succ.trans hn)), +end + +end nnreal + +open filter + +lemma filter.tendsto.nnrpow {α : Type*} {f : filter α} {u : α → ℝ≥0} {v : α → ℝ} {x : ℝ≥0} {y : ℝ} + (hx : tendsto u f (𝓝 x)) (hy : tendsto v f (𝓝 y)) (h : x ≠ 0 ∨ 0 < y) : + tendsto (λ a, (u a) ^ (v a)) f (𝓝 (x ^ y)) := +tendsto.comp (nnreal.continuous_at_rpow h) (hx.prod_mk_nhds hy) + +namespace nnreal + +lemma continuous_at_rpow_const {x : ℝ≥0} {y : ℝ} (h : x ≠ 0 ∨ 0 ≤ y) : + continuous_at (λ z, z^y) x := +h.elim (λ h, tendsto_id.nnrpow tendsto_const_nhds (or.inl h)) $ + λ h, h.eq_or_lt.elim + (λ h, h ▸ by simp only [rpow_zero, continuous_at_const]) + (λ h, tendsto_id.nnrpow tendsto_const_nhds (or.inr h)) + +lemma continuous_rpow_const {y : ℝ} (h : 0 ≤ y) : + continuous (λ x : ℝ≥0, x^y) := +continuous_iff_continuous_at.2 $ λ x, continuous_at_rpow_const (or.inr h) + +end nnreal + +/-! ## Continuity for `ℝ≥0∞` powers -/ +namespace ennreal + +lemma eventually_pow_one_div_le {x : ℝ≥0∞} (hx : x ≠ ∞) {y : ℝ≥0∞} (hy : 1 < y) : + ∀ᶠ (n : ℕ) in at_top, x ^ (1 / n : ℝ) ≤ y := +begin + lift x to ℝ≥0 using hx, + by_cases y = ∞, + { exact eventually_of_forall (λ n, h.symm ▸ le_top) }, + { lift y to ℝ≥0 using h, + have := nnreal.eventually_pow_one_div_le x (by exact_mod_cast hy : 1 < y), + refine this.congr (eventually_of_forall $ λ n, _), + rw [coe_rpow_of_nonneg x (by positivity : 0 ≤ (1 / n : ℝ)), coe_le_coe] }, +end + +private lemma continuous_at_rpow_const_of_pos {x : ℝ≥0∞} {y : ℝ} (h : 0 < y) : + continuous_at (λ a : ℝ≥0∞, a ^ y) x := +begin + by_cases hx : x = ⊤, + { rw [hx, continuous_at], + convert tendsto_rpow_at_top h, + simp [h] }, + lift x to ℝ≥0 using hx, + rw continuous_at_coe_iff, + convert continuous_coe.continuous_at.comp + (nnreal.continuous_at_rpow_const (or.inr h.le)) using 1, + ext1 x, + simp [coe_rpow_of_nonneg _ h.le] +end + +@[continuity] +lemma continuous_rpow_const {y : ℝ} : continuous (λ a : ℝ≥0∞, a ^ y) := +begin + apply continuous_iff_continuous_at.2 (λ x, _), + rcases lt_trichotomy 0 y with hy|rfl|hy, + { exact continuous_at_rpow_const_of_pos hy }, + { simp only [rpow_zero], exact continuous_at_const }, + { obtain ⟨z, hz⟩ : ∃ z, y = -z := ⟨-y, (neg_neg _).symm⟩, + have z_pos : 0 < z, by simpa [hz] using hy, + simp_rw [hz, rpow_neg], + exact continuous_inv.continuous_at.comp (continuous_at_rpow_const_of_pos z_pos) } +end + +lemma tendsto_const_mul_rpow_nhds_zero_of_pos {c : ℝ≥0∞} (hc : c ≠ ∞) {y : ℝ} (hy : 0 < y) : + tendsto (λ x : ℝ≥0∞, c * x ^ y) (𝓝 0) (𝓝 0) := +begin + convert ennreal.tendsto.const_mul (ennreal.continuous_rpow_const.tendsto 0) _, + { simp [hy] }, + { exact or.inr hc } +end + +end ennreal + +lemma filter.tendsto.ennrpow_const {α : Type*} {f : filter α} {m : α → ℝ≥0∞} {a : ℝ≥0∞} (r : ℝ) + (hm : tendsto m f (𝓝 a)) : + tendsto (λ x, (m x) ^ r) f (𝓝 (a ^ r)) := +(ennreal.continuous_rpow_const.tendsto a).comp hm diff --git a/src/analysis/special_functions/pow_deriv.lean b/src/analysis/special_functions/pow/deriv.lean similarity index 99% rename from src/analysis/special_functions/pow_deriv.lean rename to src/analysis/special_functions/pow/deriv.lean index c785353eef25c..e122335f31c1b 100644 --- a/src/analysis/special_functions/pow_deriv.lean +++ b/src/analysis/special_functions/pow/deriv.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, Sébastien Gouëzel, Rémy Degenne -/ -import analysis.special_functions.pow +import analysis.special_functions.pow.continuity import analysis.special_functions.complex.log_deriv import analysis.calculus.extend_deriv import analysis.special_functions.log.deriv diff --git a/src/analysis/special_functions/pow/nnreal.lean b/src/analysis/special_functions/pow/nnreal.lean new file mode 100644 index 0000000000000..155a69cd6c534 --- /dev/null +++ b/src/analysis/special_functions/pow/nnreal.lean @@ -0,0 +1,693 @@ +/- +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, Sébastien Gouëzel, + Rémy Degenne, David Loeffler +-/ +import analysis.special_functions.pow.real + +/-! +# Power function on `ℝ≥0` and `ℝ≥0∞` + +We construct the power functions `x ^ y` where +* `x` is a nonnegative real number and `y` is a real number; +* `x` is a number from `[0, +∞]` (a.k.a. `ℝ≥0∞`) and `y` is a real number. + +We also prove basic properties of these functions. +-/ + +noncomputable theory + +open_locale classical real nnreal ennreal big_operators complex_conjugate +open finset set + +namespace nnreal + +/-- The nonnegative real power function `x^y`, defined for `x : ℝ≥0` and `y : ℝ ` as the +restriction of the real power function. For `x > 0`, it is equal to `exp (y log x)`. For `x = 0`, +one sets `0 ^ 0 = 1` and `0 ^ y = 0` for `y ≠ 0`. -/ +noncomputable def rpow (x : ℝ≥0) (y : ℝ) : ℝ≥0 := +⟨(x : ℝ) ^ y, real.rpow_nonneg_of_nonneg x.2 y⟩ + +noncomputable instance : has_pow ℝ≥0 ℝ := ⟨rpow⟩ + +@[simp] lemma rpow_eq_pow (x : ℝ≥0) (y : ℝ) : rpow x y = x ^ y := rfl + +@[simp, norm_cast] lemma coe_rpow (x : ℝ≥0) (y : ℝ) : ((x ^ y : ℝ≥0) : ℝ) = (x : ℝ) ^ y := rfl + +@[simp] lemma rpow_zero (x : ℝ≥0) : x ^ (0 : ℝ) = 1 := +nnreal.eq $ real.rpow_zero _ + +@[simp] lemma rpow_eq_zero_iff {x : ℝ≥0} {y : ℝ} : x ^ y = 0 ↔ x = 0 ∧ y ≠ 0 := +begin + rw [← nnreal.coe_eq, coe_rpow, ← nnreal.coe_eq_zero], + exact real.rpow_eq_zero_iff_of_nonneg x.2 +end + +@[simp] lemma zero_rpow {x : ℝ} (h : x ≠ 0) : (0 : ℝ≥0) ^ x = 0 := +nnreal.eq $ real.zero_rpow h + +@[simp] lemma rpow_one (x : ℝ≥0) : x ^ (1 : ℝ) = x := +nnreal.eq $ real.rpow_one _ + +@[simp] lemma one_rpow (x : ℝ) : (1 : ℝ≥0) ^ x = 1 := +nnreal.eq $ real.one_rpow _ + +lemma rpow_add {x : ℝ≥0} (hx : x ≠ 0) (y z : ℝ) : x ^ (y + z) = x ^ y * x ^ z := +nnreal.eq $ real.rpow_add (pos_iff_ne_zero.2 hx) _ _ + +lemma rpow_add' (x : ℝ≥0) {y z : ℝ} (h : y + z ≠ 0) : x ^ (y + z) = x ^ y * x ^ z := +nnreal.eq $ real.rpow_add' x.2 h + +lemma rpow_mul (x : ℝ≥0) (y z : ℝ) : x ^ (y * z) = (x ^ y) ^ z := +nnreal.eq $ real.rpow_mul x.2 y z + +lemma rpow_neg (x : ℝ≥0) (y : ℝ) : x ^ -y = (x ^ y)⁻¹ := +nnreal.eq $ real.rpow_neg x.2 _ + +lemma rpow_neg_one (x : ℝ≥0) : x ^ (-1 : ℝ) = x ⁻¹ := +by simp [rpow_neg] + +lemma rpow_sub {x : ℝ≥0} (hx : x ≠ 0) (y z : ℝ) : x ^ (y - z) = x ^ y / x ^ z := +nnreal.eq $ real.rpow_sub (pos_iff_ne_zero.2 hx) y z + +lemma rpow_sub' (x : ℝ≥0) {y z : ℝ} (h : y - z ≠ 0) : + x ^ (y - z) = x ^ y / x ^ z := +nnreal.eq $ real.rpow_sub' x.2 h + +lemma rpow_inv_rpow_self {y : ℝ} (hy : y ≠ 0) (x : ℝ≥0) : (x ^ y) ^ (1 / y) = x := +by field_simp [← rpow_mul] + +lemma rpow_self_rpow_inv {y : ℝ} (hy : y ≠ 0) (x : ℝ≥0) : (x ^ (1 / y)) ^ y = x := +by field_simp [← rpow_mul] + +lemma inv_rpow (x : ℝ≥0) (y : ℝ) : (x⁻¹) ^ y = (x ^ y)⁻¹ := +nnreal.eq $ real.inv_rpow x.2 y + +lemma div_rpow (x y : ℝ≥0) (z : ℝ) : (x / y) ^ z = x ^ z / y ^ z := +nnreal.eq $ real.div_rpow x.2 y.2 z + +lemma sqrt_eq_rpow (x : ℝ≥0) : sqrt x = x ^ (1/(2:ℝ)) := +begin + refine nnreal.eq _, + push_cast, + exact real.sqrt_eq_rpow x.1, +end + +@[simp, norm_cast] lemma rpow_nat_cast (x : ℝ≥0) (n : ℕ) : x ^ (n : ℝ) = x ^ n := +nnreal.eq $ by simpa only [coe_rpow, coe_pow] using real.rpow_nat_cast x n + +@[simp] lemma rpow_two (x : ℝ≥0) : x ^ (2 : ℝ) = x ^ 2 := +by { rw ← rpow_nat_cast, simp only [nat.cast_bit0, nat.cast_one] } + +lemma mul_rpow {x y : ℝ≥0} {z : ℝ} : (x*y)^z = x^z * y^z := +nnreal.eq $ real.mul_rpow x.2 y.2 + +lemma rpow_le_rpow {x y : ℝ≥0} {z: ℝ} (h₁ : x ≤ y) (h₂ : 0 ≤ z) : x^z ≤ y^z := +real.rpow_le_rpow x.2 h₁ h₂ + +lemma rpow_lt_rpow {x y : ℝ≥0} {z: ℝ} (h₁ : x < y) (h₂ : 0 < z) : x^z < y^z := +real.rpow_lt_rpow x.2 h₁ h₂ + +lemma rpow_lt_rpow_iff {x y : ℝ≥0} {z : ℝ} (hz : 0 < z) : x ^ z < y ^ z ↔ x < y := +real.rpow_lt_rpow_iff x.2 y.2 hz + +lemma rpow_le_rpow_iff {x y : ℝ≥0} {z : ℝ} (hz : 0 < z) : x ^ z ≤ y ^ z ↔ x ≤ y := +real.rpow_le_rpow_iff x.2 y.2 hz + +lemma le_rpow_one_div_iff {x y : ℝ≥0} {z : ℝ} (hz : 0 < z) : x ≤ y ^ (1 / z) ↔ x ^ z ≤ y := +by rw [← rpow_le_rpow_iff hz, rpow_self_rpow_inv hz.ne'] + +lemma rpow_one_div_le_iff {x y : ℝ≥0} {z : ℝ} (hz : 0 < z) : x ^ (1 / z) ≤ y ↔ x ≤ y ^ z := +by rw [← rpow_le_rpow_iff hz, rpow_self_rpow_inv hz.ne'] + +lemma rpow_lt_rpow_of_exponent_lt {x : ℝ≥0} {y z : ℝ} (hx : 1 < x) (hyz : y < z) : x^y < x^z := +real.rpow_lt_rpow_of_exponent_lt hx hyz + +lemma rpow_le_rpow_of_exponent_le {x : ℝ≥0} {y z : ℝ} (hx : 1 ≤ x) (hyz : y ≤ z) : x^y ≤ x^z := +real.rpow_le_rpow_of_exponent_le hx hyz + +lemma rpow_lt_rpow_of_exponent_gt {x : ℝ≥0} {y z : ℝ} (hx0 : 0 < x) (hx1 : x < 1) (hyz : z < y) : + x^y < x^z := +real.rpow_lt_rpow_of_exponent_gt hx0 hx1 hyz + +lemma rpow_le_rpow_of_exponent_ge {x : ℝ≥0} {y z : ℝ} (hx0 : 0 < x) (hx1 : x ≤ 1) (hyz : z ≤ y) : + x^y ≤ x^z := +real.rpow_le_rpow_of_exponent_ge hx0 hx1 hyz + +lemma rpow_pos {p : ℝ} {x : ℝ≥0} (hx_pos : 0 < x) : 0 < x^p := +begin + have rpow_pos_of_nonneg : ∀ {p : ℝ}, 0 < p → 0 < x^p, + { intros p hp_pos, + rw ←zero_rpow hp_pos.ne', + exact rpow_lt_rpow hx_pos hp_pos }, + rcases lt_trichotomy 0 p with hp_pos|rfl|hp_neg, + { exact rpow_pos_of_nonneg hp_pos }, + { simp only [zero_lt_one, rpow_zero] }, + { rw [←neg_neg p, rpow_neg, inv_pos], + exact rpow_pos_of_nonneg (neg_pos.mpr hp_neg) }, +end + +lemma rpow_lt_one {x : ℝ≥0} {z : ℝ} (hx1 : x < 1) (hz : 0 < z) : x^z < 1 := +real.rpow_lt_one (coe_nonneg x) hx1 hz + +lemma rpow_le_one {x : ℝ≥0} {z : ℝ} (hx2 : x ≤ 1) (hz : 0 ≤ z) : x^z ≤ 1 := +real.rpow_le_one x.2 hx2 hz + +lemma rpow_lt_one_of_one_lt_of_neg {x : ℝ≥0} {z : ℝ} (hx : 1 < x) (hz : z < 0) : x^z < 1 := +real.rpow_lt_one_of_one_lt_of_neg hx hz + +lemma rpow_le_one_of_one_le_of_nonpos {x : ℝ≥0} {z : ℝ} (hx : 1 ≤ x) (hz : z ≤ 0) : x^z ≤ 1 := +real.rpow_le_one_of_one_le_of_nonpos hx hz + +lemma one_lt_rpow {x : ℝ≥0} {z : ℝ} (hx : 1 < x) (hz : 0 < z) : 1 < x^z := +real.one_lt_rpow hx hz + +lemma one_le_rpow {x : ℝ≥0} {z : ℝ} (h : 1 ≤ x) (h₁ : 0 ≤ z) : 1 ≤ x^z := +real.one_le_rpow h h₁ + +lemma one_lt_rpow_of_pos_of_lt_one_of_neg {x : ℝ≥0} {z : ℝ} (hx1 : 0 < x) (hx2 : x < 1) + (hz : z < 0) : 1 < x^z := +real.one_lt_rpow_of_pos_of_lt_one_of_neg hx1 hx2 hz + +lemma one_le_rpow_of_pos_of_le_one_of_nonpos {x : ℝ≥0} {z : ℝ} (hx1 : 0 < x) (hx2 : x ≤ 1) + (hz : z ≤ 0) : 1 ≤ x^z := +real.one_le_rpow_of_pos_of_le_one_of_nonpos hx1 hx2 hz + +lemma rpow_le_self_of_le_one {x : ℝ≥0} {z : ℝ} (hx : x ≤ 1) (h_one_le : 1 ≤ z) : x ^ z ≤ x := +begin + rcases eq_bot_or_bot_lt x with rfl | (h : 0 < x), + { have : z ≠ 0 := by linarith, + simp [this] }, + nth_rewrite 1 ←nnreal.rpow_one x, + exact nnreal.rpow_le_rpow_of_exponent_ge h hx h_one_le, +end + +lemma rpow_left_injective {x : ℝ} (hx : x ≠ 0) : function.injective (λ y : ℝ≥0, y^x) := +λ y z hyz, by simpa only [rpow_inv_rpow_self hx] using congr_arg (λ y, y ^ (1 / x)) hyz + +lemma rpow_eq_rpow_iff {x y : ℝ≥0} {z : ℝ} (hz : z ≠ 0) : x ^ z = y ^ z ↔ x = y := +(rpow_left_injective hz).eq_iff + +lemma rpow_left_surjective {x : ℝ} (hx : x ≠ 0) : function.surjective (λ y : ℝ≥0, y^x) := +λ y, ⟨y ^ x⁻¹, by simp_rw [←rpow_mul, _root_.inv_mul_cancel hx, rpow_one]⟩ + +lemma rpow_left_bijective {x : ℝ} (hx : x ≠ 0) : function.bijective (λ y : ℝ≥0, y^x) := +⟨rpow_left_injective hx, rpow_left_surjective hx⟩ + +lemma eq_rpow_one_div_iff {x y : ℝ≥0} {z : ℝ} (hz : z ≠ 0) : x = y ^ (1 / z) ↔ x ^ z = y := +by rw [← rpow_eq_rpow_iff hz, rpow_self_rpow_inv hz] + +lemma rpow_one_div_eq_iff {x y : ℝ≥0} {z : ℝ} (hz : z ≠ 0) : x ^ (1 / z) = y ↔ x = y ^ z := +by rw [← rpow_eq_rpow_iff hz, rpow_self_rpow_inv hz] + +lemma pow_nat_rpow_nat_inv (x : ℝ≥0) {n : ℕ} (hn : n ≠ 0) : + (x ^ n) ^ (n⁻¹ : ℝ) = x := +by { rw [← nnreal.coe_eq, coe_rpow, nnreal.coe_pow], exact real.pow_nat_rpow_nat_inv x.2 hn } + +lemma rpow_nat_inv_pow_nat (x : ℝ≥0) {n : ℕ} (hn : n ≠ 0) : + (x ^ (n⁻¹ : ℝ)) ^ n = x := +by { rw [← nnreal.coe_eq, nnreal.coe_pow, coe_rpow], exact real.rpow_nat_inv_pow_nat x.2 hn } + +lemma _root_.real.to_nnreal_rpow_of_nonneg {x y : ℝ} (hx : 0 ≤ x) : + real.to_nnreal (x ^ y) = (real.to_nnreal x) ^ y := +begin + nth_rewrite 0 ← real.coe_to_nnreal x hx, + rw [←nnreal.coe_rpow, real.to_nnreal_coe], +end + +end nnreal + +namespace ennreal + +/-- The real power function `x^y` on extended nonnegative reals, defined for `x : ℝ≥0∞` and +`y : ℝ` as the restriction of the real power function if `0 < x < ⊤`, and with the natural values +for `0` and `⊤` (i.e., `0 ^ x = 0` for `x > 0`, `1` for `x = 0` and `⊤` for `x < 0`, and +`⊤ ^ x = 1 / 0 ^ x`). -/ +noncomputable def rpow : ℝ≥0∞ → ℝ → ℝ≥0∞ +| (some x) y := if x = 0 ∧ y < 0 then ⊤ else (x ^ y : ℝ≥0) +| none y := if 0 < y then ⊤ else if y = 0 then 1 else 0 + +noncomputable instance : has_pow ℝ≥0∞ ℝ := ⟨rpow⟩ + +@[simp] lemma rpow_eq_pow (x : ℝ≥0∞) (y : ℝ) : rpow x y = x ^ y := rfl + +@[simp] lemma rpow_zero {x : ℝ≥0∞} : x ^ (0 : ℝ) = 1 := +by cases x; { dsimp only [(^), rpow], simp [lt_irrefl] } + +lemma top_rpow_def (y : ℝ) : (⊤ : ℝ≥0∞) ^ y = if 0 < y then ⊤ else if y = 0 then 1 else 0 := +rfl + +@[simp] lemma top_rpow_of_pos {y : ℝ} (h : 0 < y) : (⊤ : ℝ≥0∞) ^ y = ⊤ := +by simp [top_rpow_def, h] + +@[simp] lemma top_rpow_of_neg {y : ℝ} (h : y < 0) : (⊤ : ℝ≥0∞) ^ y = 0 := +by simp [top_rpow_def, asymm h, ne_of_lt h] + +@[simp] lemma zero_rpow_of_pos {y : ℝ} (h : 0 < y) : (0 : ℝ≥0∞) ^ y = 0 := +begin + rw [← ennreal.coe_zero, ← ennreal.some_eq_coe], + dsimp only [(^), rpow], + simp [h, asymm h, ne_of_gt h], +end + +@[simp] lemma zero_rpow_of_neg {y : ℝ} (h : y < 0) : (0 : ℝ≥0∞) ^ y = ⊤ := +begin + rw [← ennreal.coe_zero, ← ennreal.some_eq_coe], + dsimp only [(^), rpow], + simp [h, ne_of_gt h], +end + +lemma zero_rpow_def (y : ℝ) : (0 : ℝ≥0∞) ^ y = if 0 < y then 0 else if y = 0 then 1 else ⊤ := +begin + rcases lt_trichotomy 0 y with H|rfl|H, + { simp [H, ne_of_gt, zero_rpow_of_pos, lt_irrefl] }, + { simp [lt_irrefl] }, + { simp [H, asymm H, ne_of_lt, zero_rpow_of_neg] } +end + +@[simp] lemma zero_rpow_mul_self (y : ℝ) : (0 : ℝ≥0∞) ^ y * 0 ^ y = 0 ^ y := +by { rw zero_rpow_def, split_ifs, exacts [zero_mul _, one_mul _, top_mul_top] } + +@[norm_cast] lemma coe_rpow_of_ne_zero {x : ℝ≥0} (h : x ≠ 0) (y : ℝ) : + (x : ℝ≥0∞) ^ y = (x ^ y : ℝ≥0) := +begin + rw [← ennreal.some_eq_coe], + dsimp only [(^), rpow], + simp [h] +end + +@[norm_cast] lemma coe_rpow_of_nonneg (x : ℝ≥0) {y : ℝ} (h : 0 ≤ y) : + (x : ℝ≥0∞) ^ y = (x ^ y : ℝ≥0) := +begin + by_cases hx : x = 0, + { rcases le_iff_eq_or_lt.1 h with H|H, + { simp [hx, H.symm] }, + { simp [hx, zero_rpow_of_pos H, nnreal.zero_rpow (ne_of_gt H)] } }, + { exact coe_rpow_of_ne_zero hx _ } +end + +lemma coe_rpow_def (x : ℝ≥0) (y : ℝ) : + (x : ℝ≥0∞) ^ y = if x = 0 ∧ y < 0 then ⊤ else (x ^ y : ℝ≥0) := rfl + +@[simp] lemma rpow_one (x : ℝ≥0∞) : x ^ (1 : ℝ) = x := +begin + cases x, + { exact dif_pos zero_lt_one }, + { change ite _ _ _ = _, + simp only [nnreal.rpow_one, some_eq_coe, ite_eq_right_iff, top_ne_coe, and_imp], + exact λ _, zero_le_one.not_lt } +end + +@[simp] lemma one_rpow (x : ℝ) : (1 : ℝ≥0∞) ^ x = 1 := +by { rw [← coe_one, coe_rpow_of_ne_zero one_ne_zero], simp } + +@[simp] lemma rpow_eq_zero_iff {x : ℝ≥0∞} {y : ℝ} : + x ^ y = 0 ↔ (x = 0 ∧ 0 < y) ∨ (x = ⊤ ∧ y < 0) := +begin + cases x, + { rcases lt_trichotomy y 0 with H|H|H; + simp [H, top_rpow_of_neg, top_rpow_of_pos, le_of_lt] }, + { by_cases h : x = 0, + { rcases lt_trichotomy y 0 with H|H|H; + simp [h, H, zero_rpow_of_neg, zero_rpow_of_pos, le_of_lt] }, + { simp [coe_rpow_of_ne_zero h, h] } } +end + +@[simp] lemma rpow_eq_top_iff {x : ℝ≥0∞} {y : ℝ} : + x ^ y = ⊤ ↔ (x = 0 ∧ y < 0) ∨ (x = ⊤ ∧ 0 < y) := +begin + cases x, + { rcases lt_trichotomy y 0 with H|H|H; + simp [H, top_rpow_of_neg, top_rpow_of_pos, le_of_lt] }, + { by_cases h : x = 0, + { rcases lt_trichotomy y 0 with H|H|H; + simp [h, H, zero_rpow_of_neg, zero_rpow_of_pos, le_of_lt] }, + { simp [coe_rpow_of_ne_zero h, h] } } +end + +lemma rpow_eq_top_iff_of_pos {x : ℝ≥0∞} {y : ℝ} (hy : 0 < y) : x ^ y = ⊤ ↔ x = ⊤ := +by simp [rpow_eq_top_iff, hy, asymm hy] + +lemma rpow_eq_top_of_nonneg (x : ℝ≥0∞) {y : ℝ} (hy0 : 0 ≤ y) : x ^ y = ⊤ → x = ⊤ := +begin + rw ennreal.rpow_eq_top_iff, + intro h, + cases h, + { exfalso, rw lt_iff_not_ge at h, exact h.right hy0, }, + { exact h.left, }, +end + +lemma rpow_ne_top_of_nonneg {x : ℝ≥0∞} {y : ℝ} (hy0 : 0 ≤ y) (h : x ≠ ⊤) : x ^ y ≠ ⊤ := +mt (ennreal.rpow_eq_top_of_nonneg x hy0) h + +lemma rpow_lt_top_of_nonneg {x : ℝ≥0∞} {y : ℝ} (hy0 : 0 ≤ y) (h : x ≠ ⊤) : x ^ y < ⊤ := +lt_top_iff_ne_top.mpr (ennreal.rpow_ne_top_of_nonneg hy0 h) + +lemma rpow_add {x : ℝ≥0∞} (y z : ℝ) (hx : x ≠ 0) (h'x : x ≠ ⊤) : x ^ (y + z) = x ^ y * x ^ z := +begin + cases x, { exact (h'x rfl).elim }, + have : x ≠ 0 := λ h, by simpa [h] using hx, + simp [coe_rpow_of_ne_zero this, nnreal.rpow_add this] +end + +lemma rpow_neg (x : ℝ≥0∞) (y : ℝ) : x ^ -y = (x ^ y)⁻¹ := +begin + cases x, + { rcases lt_trichotomy y 0 with H|H|H; + simp [top_rpow_of_pos, top_rpow_of_neg, H, neg_pos.mpr] }, + { by_cases h : x = 0, + { rcases lt_trichotomy y 0 with H|H|H; + simp [h, zero_rpow_of_pos, zero_rpow_of_neg, H, neg_pos.mpr] }, + { have A : x ^ y ≠ 0, by simp [h], + simp [coe_rpow_of_ne_zero h, ← coe_inv A, nnreal.rpow_neg] } } +end + +lemma rpow_sub {x : ℝ≥0∞} (y z : ℝ) (hx : x ≠ 0) (h'x : x ≠ ⊤) : x ^ (y - z) = x ^ y / x ^ z := +by rw [sub_eq_add_neg, rpow_add _ _ hx h'x, rpow_neg, div_eq_mul_inv] + +lemma rpow_neg_one (x : ℝ≥0∞) : x ^ (-1 : ℝ) = x ⁻¹ := +by simp [rpow_neg] + +lemma rpow_mul (x : ℝ≥0∞) (y z : ℝ) : x ^ (y * z) = (x ^ y) ^ z := +begin + cases x, + { rcases lt_trichotomy y 0 with Hy|Hy|Hy; + rcases lt_trichotomy z 0 with Hz|Hz|Hz; + simp [Hy, Hz, zero_rpow_of_neg, zero_rpow_of_pos, top_rpow_of_neg, top_rpow_of_pos, + mul_pos_of_neg_of_neg, mul_neg_of_neg_of_pos, mul_neg_of_pos_of_neg] }, + { by_cases h : x = 0, + { rcases lt_trichotomy y 0 with Hy|Hy|Hy; + rcases lt_trichotomy z 0 with Hz|Hz|Hz; + simp [h, Hy, Hz, zero_rpow_of_neg, zero_rpow_of_pos, top_rpow_of_neg, top_rpow_of_pos, + mul_pos_of_neg_of_neg, mul_neg_of_neg_of_pos, mul_neg_of_pos_of_neg] }, + { have : x ^ y ≠ 0, by simp [h], + simp [coe_rpow_of_ne_zero h, coe_rpow_of_ne_zero this, nnreal.rpow_mul] } } +end + +@[simp, norm_cast] lemma rpow_nat_cast (x : ℝ≥0∞) (n : ℕ) : x ^ (n : ℝ) = x ^ n := +begin + cases x, + { cases n; + simp [top_rpow_of_pos (nat.cast_add_one_pos _), top_pow (nat.succ_pos _)] }, + { simp [coe_rpow_of_nonneg _ (nat.cast_nonneg n)] } +end + +@[simp] lemma rpow_two (x : ℝ≥0∞) : x ^ (2 : ℝ) = x ^ 2 := +by { rw ← rpow_nat_cast, simp only [nat.cast_bit0, nat.cast_one] } + +lemma mul_rpow_eq_ite (x y : ℝ≥0∞) (z : ℝ) : + (x * y) ^ z = if (x = 0 ∧ y = ⊤ ∨ x = ⊤ ∧ y = 0) ∧ z < 0 then ⊤ else x ^ z * y ^ z := +begin + rcases eq_or_ne z 0 with rfl|hz, { simp }, + replace hz := hz.lt_or_lt, + wlog hxy : x ≤ y, + { convert this y x z hz (le_of_not_le hxy) using 2; simp only [mul_comm, and_comm, or_comm], }, + rcases eq_or_ne x 0 with rfl|hx0, + { induction y using with_top.rec_top_coe; cases hz with hz hz; simp [*, hz.not_lt] }, + rcases eq_or_ne y 0 with rfl|hy0, { exact (hx0 (bot_unique hxy)).elim }, + induction x using with_top.rec_top_coe, { cases hz with hz hz; simp [hz, top_unique hxy] }, + induction y using with_top.rec_top_coe, { cases hz with hz hz; simp * }, + simp only [*, false_and, and_false, false_or, if_false], + norm_cast at *, + rw [coe_rpow_of_ne_zero (mul_ne_zero hx0 hy0), nnreal.mul_rpow] +end + +lemma mul_rpow_of_ne_top {x y : ℝ≥0∞} (hx : x ≠ ⊤) (hy : y ≠ ⊤) (z : ℝ) : + (x * y) ^ z = x^z * y^z := +by simp [*, mul_rpow_eq_ite] + +@[norm_cast] lemma coe_mul_rpow (x y : ℝ≥0) (z : ℝ) : + ((x : ℝ≥0∞) * y) ^ z = x^z * y^z := +mul_rpow_of_ne_top coe_ne_top coe_ne_top z + +lemma mul_rpow_of_ne_zero {x y : ℝ≥0∞} (hx : x ≠ 0) (hy : y ≠ 0) (z : ℝ) : + (x * y) ^ z = x ^ z * y ^ z := +by simp [*, mul_rpow_eq_ite] + +lemma mul_rpow_of_nonneg (x y : ℝ≥0∞) {z : ℝ} (hz : 0 ≤ z) : + (x * y) ^ z = x ^ z * y ^ z := +by simp [hz.not_lt, mul_rpow_eq_ite] + +lemma inv_rpow (x : ℝ≥0∞) (y : ℝ) : (x⁻¹) ^ y = (x ^ y)⁻¹ := +begin + rcases eq_or_ne y 0 with rfl|hy, { simp only [rpow_zero, inv_one] }, + replace hy := hy.lt_or_lt, + rcases eq_or_ne x 0 with rfl|h0, { cases hy; simp * }, + rcases eq_or_ne x ⊤ with rfl|h_top, { cases hy; simp * }, + apply ennreal.eq_inv_of_mul_eq_one_left, + rw [← mul_rpow_of_ne_zero (ennreal.inv_ne_zero.2 h_top) h0, ennreal.inv_mul_cancel h0 h_top, + one_rpow] +end + +lemma div_rpow_of_nonneg (x y : ℝ≥0∞) {z : ℝ} (hz : 0 ≤ z) : + (x / y) ^ z = x ^ z / y ^ z := +by rw [div_eq_mul_inv, mul_rpow_of_nonneg _ _ hz, inv_rpow, div_eq_mul_inv] + +lemma strict_mono_rpow_of_pos {z : ℝ} (h : 0 < z) : strict_mono (λ x : ℝ≥0∞, x ^ z) := +begin + intros x y hxy, + lift x to ℝ≥0 using ne_top_of_lt hxy, + rcases eq_or_ne y ∞ with rfl|hy, + { simp only [top_rpow_of_pos h, coe_rpow_of_nonneg _ h.le, coe_lt_top] }, + { lift y to ℝ≥0 using hy, + simp only [coe_rpow_of_nonneg _ h.le, nnreal.rpow_lt_rpow (coe_lt_coe.1 hxy) h, coe_lt_coe] } +end + +lemma monotone_rpow_of_nonneg {z : ℝ} (h : 0 ≤ z) : monotone (λ x : ℝ≥0∞, x ^ z) := +h.eq_or_lt.elim (λ h0, h0 ▸ by simp only [rpow_zero, monotone_const]) + (λ h0, (strict_mono_rpow_of_pos h0).monotone) + +/-- Bundles `λ x : ℝ≥0∞, x ^ y` into an order isomorphism when `y : ℝ` is positive, +where the inverse is `λ x : ℝ≥0∞, x ^ (1 / y)`. -/ +@[simps apply] def order_iso_rpow (y : ℝ) (hy : 0 < y) : ℝ≥0∞ ≃o ℝ≥0∞ := +(strict_mono_rpow_of_pos hy).order_iso_of_right_inverse (λ x, x ^ y) (λ x, x ^ (1 / y)) + (λ x, by { dsimp, rw [←rpow_mul, one_div_mul_cancel hy.ne.symm, rpow_one] }) + +lemma order_iso_rpow_symm_apply (y : ℝ) (hy : 0 < y) : + (order_iso_rpow y hy).symm = order_iso_rpow (1 / y) (one_div_pos.2 hy) := +by { simp only [order_iso_rpow, one_div_one_div], refl } + +lemma rpow_le_rpow {x y : ℝ≥0∞} {z : ℝ} (h₁ : x ≤ y) (h₂ : 0 ≤ z) : x^z ≤ y^z := +monotone_rpow_of_nonneg h₂ h₁ + +lemma rpow_lt_rpow {x y : ℝ≥0∞} {z : ℝ} (h₁ : x < y) (h₂ : 0 < z) : x^z < y^z := +strict_mono_rpow_of_pos h₂ h₁ + +lemma rpow_le_rpow_iff {x y : ℝ≥0∞} {z : ℝ} (hz : 0 < z) : x ^ z ≤ y ^ z ↔ x ≤ y := +(strict_mono_rpow_of_pos hz).le_iff_le + +lemma rpow_lt_rpow_iff {x y : ℝ≥0∞} {z : ℝ} (hz : 0 < z) : x ^ z < y ^ z ↔ x < y := +(strict_mono_rpow_of_pos hz).lt_iff_lt + +lemma le_rpow_one_div_iff {x y : ℝ≥0∞} {z : ℝ} (hz : 0 < z) : x ≤ y ^ (1 / z) ↔ x ^ z ≤ y := +begin + nth_rewrite 0 ←rpow_one x, + nth_rewrite 0 ←@_root_.mul_inv_cancel _ _ z hz.ne', + rw [rpow_mul, ←one_div, @rpow_le_rpow_iff _ _ (1/z) (by simp [hz])], +end + +lemma lt_rpow_one_div_iff {x y : ℝ≥0∞} {z : ℝ} (hz : 0 < z) : x < y ^ (1 / z) ↔ x ^ z < y := +begin + nth_rewrite 0 ←rpow_one x, + nth_rewrite 0 ←@_root_.mul_inv_cancel _ _ z (ne_of_lt hz).symm, + rw [rpow_mul, ←one_div, @rpow_lt_rpow_iff _ _ (1/z) (by simp [hz])], +end + +lemma rpow_one_div_le_iff {x y : ℝ≥0∞} {z : ℝ} (hz : 0 < z) : x ^ (1 / z) ≤ y ↔ x ≤ y ^ z := +begin + nth_rewrite 0 ← ennreal.rpow_one y, + nth_rewrite 1 ← @_root_.mul_inv_cancel _ _ z hz.ne.symm, + rw [ennreal.rpow_mul, ← one_div, ennreal.rpow_le_rpow_iff (one_div_pos.2 hz)], +end + +lemma rpow_lt_rpow_of_exponent_lt {x : ℝ≥0∞} {y z : ℝ} (hx : 1 < x) (hx' : x ≠ ⊤) (hyz : y < z) : + x^y < x^z := +begin + lift x to ℝ≥0 using hx', + rw [one_lt_coe_iff] at hx, + simp [coe_rpow_of_ne_zero (ne_of_gt (lt_trans zero_lt_one hx)), + nnreal.rpow_lt_rpow_of_exponent_lt hx hyz] +end + +lemma rpow_le_rpow_of_exponent_le {x : ℝ≥0∞} {y z : ℝ} (hx : 1 ≤ x) (hyz : y ≤ z) : x^y ≤ x^z := +begin + cases x, + { rcases lt_trichotomy y 0 with Hy|Hy|Hy; + rcases lt_trichotomy z 0 with Hz|Hz|Hz; + simp [Hy, Hz, top_rpow_of_neg, top_rpow_of_pos, le_refl]; + linarith }, + { simp only [one_le_coe_iff, some_eq_coe] at hx, + simp [coe_rpow_of_ne_zero (ne_of_gt (lt_of_lt_of_le zero_lt_one hx)), + nnreal.rpow_le_rpow_of_exponent_le hx hyz] } +end + +lemma rpow_lt_rpow_of_exponent_gt {x : ℝ≥0∞} {y z : ℝ} (hx0 : 0 < x) (hx1 : x < 1) (hyz : z < y) : + x^y < x^z := +begin + lift x to ℝ≥0 using ne_of_lt (lt_of_lt_of_le hx1 le_top), + simp only [coe_lt_one_iff, coe_pos] at hx0 hx1, + simp [coe_rpow_of_ne_zero (ne_of_gt hx0), nnreal.rpow_lt_rpow_of_exponent_gt hx0 hx1 hyz] +end + +lemma rpow_le_rpow_of_exponent_ge {x : ℝ≥0∞} {y z : ℝ} (hx1 : x ≤ 1) (hyz : z ≤ y) : + x^y ≤ x^z := +begin + lift x to ℝ≥0 using ne_of_lt (lt_of_le_of_lt hx1 coe_lt_top), + by_cases h : x = 0, + { rcases lt_trichotomy y 0 with Hy|Hy|Hy; + rcases lt_trichotomy z 0 with Hz|Hz|Hz; + simp [Hy, Hz, h, zero_rpow_of_neg, zero_rpow_of_pos, le_refl]; + linarith }, + { rw [coe_le_one_iff] at hx1, + simp [coe_rpow_of_ne_zero h, + nnreal.rpow_le_rpow_of_exponent_ge (bot_lt_iff_ne_bot.mpr h) hx1 hyz] } +end + +lemma rpow_le_self_of_le_one {x : ℝ≥0∞} {z : ℝ} (hx : x ≤ 1) (h_one_le : 1 ≤ z) : x ^ z ≤ x := +begin + nth_rewrite 1 ←ennreal.rpow_one x, + exact ennreal.rpow_le_rpow_of_exponent_ge hx h_one_le, +end + +lemma le_rpow_self_of_one_le {x : ℝ≥0∞} {z : ℝ} (hx : 1 ≤ x) (h_one_le : 1 ≤ z) : x ≤ x ^ z := +begin + nth_rewrite 0 ←ennreal.rpow_one x, + exact ennreal.rpow_le_rpow_of_exponent_le hx h_one_le, +end + +lemma rpow_pos_of_nonneg {p : ℝ} {x : ℝ≥0∞} (hx_pos : 0 < x) (hp_nonneg : 0 ≤ p) : 0 < x^p := +begin + by_cases hp_zero : p = 0, + { simp [hp_zero, zero_lt_one], }, + { rw ←ne.def at hp_zero, + have hp_pos := lt_of_le_of_ne hp_nonneg hp_zero.symm, + rw ←zero_rpow_of_pos hp_pos, exact rpow_lt_rpow hx_pos hp_pos, }, +end + +lemma rpow_pos {p : ℝ} {x : ℝ≥0∞} (hx_pos : 0 < x) (hx_ne_top : x ≠ ⊤) : 0 < x^p := +begin + cases lt_or_le 0 p with hp_pos hp_nonpos, + { exact rpow_pos_of_nonneg hx_pos (le_of_lt hp_pos), }, + { rw [←neg_neg p, rpow_neg, ennreal.inv_pos], + exact rpow_ne_top_of_nonneg (right.nonneg_neg_iff.mpr hp_nonpos) hx_ne_top, }, +end + +lemma rpow_lt_one {x : ℝ≥0∞} {z : ℝ} (hx : x < 1) (hz : 0 < z) : x^z < 1 := +begin + lift x to ℝ≥0 using ne_of_lt (lt_of_lt_of_le hx le_top), + simp only [coe_lt_one_iff] at hx, + simp [coe_rpow_of_nonneg _ (le_of_lt hz), nnreal.rpow_lt_one hx hz], +end + +lemma rpow_le_one {x : ℝ≥0∞} {z : ℝ} (hx : x ≤ 1) (hz : 0 ≤ z) : x^z ≤ 1 := +begin + lift x to ℝ≥0 using ne_of_lt (lt_of_le_of_lt hx coe_lt_top), + simp only [coe_le_one_iff] at hx, + simp [coe_rpow_of_nonneg _ hz, nnreal.rpow_le_one hx hz], +end + +lemma rpow_lt_one_of_one_lt_of_neg {x : ℝ≥0∞} {z : ℝ} (hx : 1 < x) (hz : z < 0) : x^z < 1 := +begin + cases x, + { simp [top_rpow_of_neg hz, zero_lt_one] }, + { simp only [some_eq_coe, one_lt_coe_iff] at hx, + simp [coe_rpow_of_ne_zero (ne_of_gt (lt_trans zero_lt_one hx)), + nnreal.rpow_lt_one_of_one_lt_of_neg hx hz] }, +end + +lemma rpow_le_one_of_one_le_of_neg {x : ℝ≥0∞} {z : ℝ} (hx : 1 ≤ x) (hz : z < 0) : x^z ≤ 1 := +begin + cases x, + { simp [top_rpow_of_neg hz, zero_lt_one] }, + { simp only [one_le_coe_iff, some_eq_coe] at hx, + simp [coe_rpow_of_ne_zero (ne_of_gt (lt_of_lt_of_le zero_lt_one hx)), + nnreal.rpow_le_one_of_one_le_of_nonpos hx (le_of_lt hz)] }, +end + +lemma one_lt_rpow {x : ℝ≥0∞} {z : ℝ} (hx : 1 < x) (hz : 0 < z) : 1 < x^z := +begin + cases x, + { simp [top_rpow_of_pos hz] }, + { simp only [some_eq_coe, one_lt_coe_iff] at hx, + simp [coe_rpow_of_nonneg _ (le_of_lt hz), nnreal.one_lt_rpow hx hz] } +end + +lemma one_le_rpow {x : ℝ≥0∞} {z : ℝ} (hx : 1 ≤ x) (hz : 0 < z) : 1 ≤ x^z := +begin + cases x, + { simp [top_rpow_of_pos hz] }, + { simp only [one_le_coe_iff, some_eq_coe] at hx, + simp [coe_rpow_of_nonneg _ (le_of_lt hz), nnreal.one_le_rpow hx (le_of_lt hz)] }, +end + +lemma one_lt_rpow_of_pos_of_lt_one_of_neg {x : ℝ≥0∞} {z : ℝ} (hx1 : 0 < x) (hx2 : x < 1) + (hz : z < 0) : 1 < x^z := +begin + lift x to ℝ≥0 using ne_of_lt (lt_of_lt_of_le hx2 le_top), + simp only [coe_lt_one_iff, coe_pos] at ⊢ hx1 hx2, + simp [coe_rpow_of_ne_zero (ne_of_gt hx1), nnreal.one_lt_rpow_of_pos_of_lt_one_of_neg hx1 hx2 hz], +end + +lemma one_le_rpow_of_pos_of_le_one_of_neg {x : ℝ≥0∞} {z : ℝ} (hx1 : 0 < x) (hx2 : x ≤ 1) + (hz : z < 0) : 1 ≤ x^z := +begin + lift x to ℝ≥0 using ne_of_lt (lt_of_le_of_lt hx2 coe_lt_top), + simp only [coe_le_one_iff, coe_pos] at ⊢ hx1 hx2, + simp [coe_rpow_of_ne_zero (ne_of_gt hx1), + nnreal.one_le_rpow_of_pos_of_le_one_of_nonpos hx1 hx2 (le_of_lt hz)], +end + +lemma to_nnreal_rpow (x : ℝ≥0∞) (z : ℝ) : (x.to_nnreal) ^ z = (x ^ z).to_nnreal := +begin + rcases lt_trichotomy z 0 with H|H|H, + { cases x, { simp [H, ne_of_lt] }, + by_cases hx : x = 0, + { simp [hx, H, ne_of_lt] }, + { simp [coe_rpow_of_ne_zero hx] } }, + { simp [H] }, + { cases x, { simp [H, ne_of_gt] }, + simp [coe_rpow_of_nonneg _ (le_of_lt H)] } +end + +lemma to_real_rpow (x : ℝ≥0∞) (z : ℝ) : (x.to_real) ^ z = (x ^ z).to_real := +by rw [ennreal.to_real, ennreal.to_real, ←nnreal.coe_rpow, ennreal.to_nnreal_rpow] + +lemma of_real_rpow_of_pos {x p : ℝ} (hx_pos : 0 < x) : + ennreal.of_real x ^ p = ennreal.of_real (x ^ p) := +begin + simp_rw ennreal.of_real, + rw [coe_rpow_of_ne_zero, coe_eq_coe, real.to_nnreal_rpow_of_nonneg hx_pos.le], + simp [hx_pos], +end + +lemma of_real_rpow_of_nonneg {x p : ℝ} (hx_nonneg : 0 ≤ x) (hp_nonneg : 0 ≤ p) : + ennreal.of_real x ^ p = ennreal.of_real (x ^ p) := +begin + by_cases hp0 : p = 0, + { simp [hp0], }, + by_cases hx0 : x = 0, + { rw ← ne.def at hp0, + have hp_pos : 0 < p := lt_of_le_of_ne hp_nonneg hp0.symm, + simp [hx0, hp_pos, hp_pos.ne.symm], }, + rw ← ne.def at hx0, + exact of_real_rpow_of_pos (hx_nonneg.lt_of_ne hx0.symm), +end + +lemma rpow_left_injective {x : ℝ} (hx : x ≠ 0) : + function.injective (λ y : ℝ≥0∞, y^x) := +begin + intros y z hyz, + dsimp only at hyz, + rw [←rpow_one y, ←rpow_one z, ←_root_.mul_inv_cancel hx, rpow_mul, rpow_mul, hyz], +end + +lemma rpow_left_surjective {x : ℝ} (hx : x ≠ 0) : + function.surjective (λ y : ℝ≥0∞, y^x) := +λ y, ⟨y ^ x⁻¹, by simp_rw [←rpow_mul, _root_.inv_mul_cancel hx, rpow_one]⟩ + +lemma rpow_left_bijective {x : ℝ} (hx : x ≠ 0) : + function.bijective (λ y : ℝ≥0∞, y^x) := +⟨rpow_left_injective hx, rpow_left_surjective hx⟩ + +end ennreal diff --git a/src/analysis/special_functions/pow/real.lean b/src/analysis/special_functions/pow/real.lean new file mode 100644 index 0000000000000..c938399127f70 --- /dev/null +++ b/src/analysis/special_functions/pow/real.lean @@ -0,0 +1,641 @@ +/- +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, Sébastien Gouëzel, + Rémy Degenne, David Loeffler +-/ +import analysis.special_functions.pow.complex + +/-! # Power function on `ℝ` + +We construct the power functions `x ^ y`, where `x` and `y` are real numbers. +-/ + +noncomputable theory + +open_locale classical real big_operators complex_conjugate +open finset set + +/- +## Definitions +-/ + +namespace real + +/-- The real power function `x ^ y`, defined as the real part of the complex power function. +For `x > 0`, it is equal to `exp (y log x)`. For `x = 0`, one sets `0 ^ 0=1` and `0 ^ y=0` for +`y ≠ 0`. For `x < 0`, the definition is somewhat arbitary as it depends on the choice of a complex +determination of the logarithm. With our conventions, it is equal to `exp (y log x) cos (π y)`. -/ +noncomputable def rpow (x y : ℝ) := ((x : ℂ) ^ (y : ℂ)).re + +noncomputable instance : has_pow ℝ ℝ := ⟨rpow⟩ + +@[simp] lemma rpow_eq_pow (x y : ℝ) : rpow x y = x ^ y := rfl + +lemma rpow_def (x y : ℝ) : x ^ y = ((x : ℂ) ^ (y : ℂ)).re := rfl + +lemma rpow_def_of_nonneg {x : ℝ} (hx : 0 ≤ x) (y : ℝ) : x ^ y = + if x = 0 + then if y = 0 + then 1 + else 0 + else exp (log x * y) := +by simp only [rpow_def, complex.cpow_def]; + split_ifs; + simp [*, (complex.of_real_log hx).symm, -complex.of_real_mul, -is_R_or_C.of_real_mul, + (complex.of_real_mul _ _).symm, complex.exp_of_real_re] at * + +lemma rpow_def_of_pos {x : ℝ} (hx : 0 < x) (y : ℝ) : x ^ y = exp (log x * y) := +by rw [rpow_def_of_nonneg (le_of_lt hx), if_neg (ne_of_gt hx)] + +lemma exp_mul (x y : ℝ) : exp (x * y) = (exp x) ^ y := +by rw [rpow_def_of_pos (exp_pos _), log_exp] + +@[simp] lemma exp_one_rpow (x : ℝ) : exp 1 ^ x = exp x := by rw [←exp_mul, one_mul] + +lemma rpow_eq_zero_iff_of_nonneg {x y : ℝ} (hx : 0 ≤ x) : x ^ y = 0 ↔ x = 0 ∧ y ≠ 0 := +by { simp only [rpow_def_of_nonneg hx], split_ifs; simp [*, exp_ne_zero] } + +open_locale real + +lemma rpow_def_of_neg {x : ℝ} (hx : x < 0) (y : ℝ) : x ^ y = exp (log x * y) * cos (y * π) := +begin + rw [rpow_def, complex.cpow_def, if_neg], + have : complex.log x * y = ↑(log(-x) * y) + ↑(y * π) * complex.I, + { simp only [complex.log, abs_of_neg hx, complex.arg_of_real_of_neg hx, + complex.abs_of_real, complex.of_real_mul], ring }, + { rw [this, complex.exp_add_mul_I, ← complex.of_real_exp, ← complex.of_real_cos, + ← complex.of_real_sin, mul_add, ← complex.of_real_mul, ← mul_assoc, ← complex.of_real_mul, + complex.add_re, complex.of_real_re, complex.mul_re, complex.I_re, complex.of_real_im, + real.log_neg_eq_log], + ring }, + { rw complex.of_real_eq_zero, exact ne_of_lt hx } +end + +lemma rpow_def_of_nonpos {x : ℝ} (hx : x ≤ 0) (y : ℝ) : x ^ y = + if x = 0 + then if y = 0 + then 1 + else 0 + else exp (log x * y) * cos (y * π) := +by split_ifs; simp [rpow_def, *]; exact rpow_def_of_neg (lt_of_le_of_ne hx h) _ + +lemma rpow_pos_of_pos {x : ℝ} (hx : 0 < x) (y : ℝ) : 0 < x ^ y := +by rw rpow_def_of_pos hx; apply exp_pos + +@[simp] lemma rpow_zero (x : ℝ) : x ^ (0 : ℝ) = 1 := by simp [rpow_def] + +@[simp] lemma zero_rpow {x : ℝ} (h : x ≠ 0) : (0 : ℝ) ^ x = 0 := +by simp [rpow_def, *] + +lemma zero_rpow_eq_iff {x : ℝ} {a : ℝ} : 0 ^ x = a ↔ (x ≠ 0 ∧ a = 0) ∨ (x = 0 ∧ a = 1) := +begin + split, + { intros hyp, + simp only [rpow_def, complex.of_real_zero] at hyp, + by_cases x = 0, + { subst h, + simp only [complex.one_re, complex.of_real_zero, complex.cpow_zero] at hyp, + exact or.inr ⟨rfl, hyp.symm⟩}, + { rw complex.zero_cpow (complex.of_real_ne_zero.mpr h) at hyp, + exact or.inl ⟨h, hyp.symm⟩, }, }, + { rintro (⟨h,rfl⟩|⟨rfl,rfl⟩), + { exact zero_rpow h, }, + { exact rpow_zero _, }, }, +end + +lemma eq_zero_rpow_iff {x : ℝ} {a : ℝ} : a = 0 ^ x ↔ (x ≠ 0 ∧ a = 0) ∨ (x = 0 ∧ a = 1) := +by rw [←zero_rpow_eq_iff, eq_comm] + +@[simp] lemma rpow_one (x : ℝ) : x ^ (1 : ℝ) = x := by simp [rpow_def] + +@[simp] lemma one_rpow (x : ℝ) : (1 : ℝ) ^ x = 1 := by simp [rpow_def] + +lemma zero_rpow_le_one (x : ℝ) : (0 : ℝ) ^ x ≤ 1 := +by { by_cases h : x = 0; simp [h, zero_le_one] } + +lemma zero_rpow_nonneg (x : ℝ) : 0 ≤ (0 : ℝ) ^ x := +by { by_cases h : x = 0; simp [h, zero_le_one] } + +lemma rpow_nonneg_of_nonneg {x : ℝ} (hx : 0 ≤ x) (y : ℝ) : 0 ≤ x ^ y := +by rw [rpow_def_of_nonneg hx]; + split_ifs; simp only [zero_le_one, le_refl, le_of_lt (exp_pos _)] + +lemma abs_rpow_of_nonneg {x y : ℝ} (hx_nonneg : 0 ≤ x) : |x ^ y| = |x| ^ y := +begin + have h_rpow_nonneg : 0 ≤ x ^ y, from real.rpow_nonneg_of_nonneg hx_nonneg _, + rw [abs_eq_self.mpr hx_nonneg, abs_eq_self.mpr h_rpow_nonneg], +end + +lemma abs_rpow_le_abs_rpow (x y : ℝ) : |x ^ y| ≤ |x| ^ y := +begin + cases le_or_lt 0 x with hx hx, + { rw [abs_rpow_of_nonneg hx] }, + { rw [abs_of_neg hx, rpow_def_of_neg hx, rpow_def_of_pos (neg_pos.2 hx), log_neg_eq_log, + abs_mul, abs_of_pos (exp_pos _)], + exact mul_le_of_le_one_right (exp_pos _).le (abs_cos_le_one _) } +end + +lemma abs_rpow_le_exp_log_mul (x y : ℝ) : |x ^ y| ≤ exp (log x * y) := +begin + refine (abs_rpow_le_abs_rpow x y).trans _, + by_cases hx : x = 0, + { by_cases hy : y = 0; simp [hx, hy, zero_le_one] }, + { rw [rpow_def_of_pos (abs_pos.2 hx), log_abs] } +end + +lemma norm_rpow_of_nonneg {x y : ℝ} (hx_nonneg : 0 ≤ x) : ‖x ^ y‖ = ‖x‖ ^ y := +by { simp_rw real.norm_eq_abs, exact abs_rpow_of_nonneg hx_nonneg, } + + +variables {x y z : ℝ} + +lemma rpow_add (hx : 0 < x) (y z : ℝ) : x ^ (y + z) = x ^ y * x ^ z := +by simp only [rpow_def_of_pos hx, mul_add, exp_add] + +lemma rpow_add' (hx : 0 ≤ x) (h : y + z ≠ 0) : x ^ (y + z) = x ^ y * x ^ z := +begin + rcases hx.eq_or_lt with rfl|pos, + { rw [zero_rpow h, zero_eq_mul], + have : y ≠ 0 ∨ z ≠ 0, from not_and_distrib.1 (λ ⟨hy, hz⟩, h $ hy.symm ▸ hz.symm ▸ zero_add 0), + exact this.imp zero_rpow zero_rpow }, + { exact rpow_add pos _ _ } +end + +lemma rpow_add_of_nonneg (hx : 0 ≤ x) (hy : 0 ≤ y) (hz : 0 ≤ z) : + x ^ (y + z) = x ^ y * x ^ z := +begin + rcases hy.eq_or_lt with rfl|hy, + { rw [zero_add, rpow_zero, one_mul] }, + exact rpow_add' hx (ne_of_gt $ add_pos_of_pos_of_nonneg hy hz) +end + +/-- For `0 ≤ x`, the only problematic case in the equality `x ^ y * x ^ z = x ^ (y + z)` is for +`x = 0` and `y + z = 0`, where the right hand side is `1` while the left hand side can vanish. +The inequality is always true, though, and given in this lemma. -/ +lemma le_rpow_add {x : ℝ} (hx : 0 ≤ x) (y z : ℝ) : x ^ y * x ^ z ≤ x ^ (y + z) := +begin + rcases le_iff_eq_or_lt.1 hx with H|pos, + { by_cases h : y + z = 0, + { simp only [H.symm, h, rpow_zero], + calc (0 : ℝ) ^ y * 0 ^ z ≤ 1 * 1 : + mul_le_mul (zero_rpow_le_one y) (zero_rpow_le_one z) (zero_rpow_nonneg z) zero_le_one + ... = 1 : by simp }, + { simp [rpow_add', ← H, h] } }, + { simp [rpow_add pos] } +end + +lemma rpow_sum_of_pos {ι : Type*} {a : ℝ} (ha : 0 < a) (f : ι → ℝ) (s : finset ι) : + a ^ (∑ x in s, f x) = ∏ x in s, a ^ f x := +@add_monoid_hom.map_sum ℝ ι (additive ℝ) _ _ ⟨λ x : ℝ, (a ^ x : ℝ), rpow_zero a, rpow_add ha⟩ f s + +lemma rpow_sum_of_nonneg {ι : Type*} {a : ℝ} (ha : 0 ≤ a) {s : finset ι} {f : ι → ℝ} + (h : ∀ x ∈ s, 0 ≤ f x) : + a ^ (∑ x in s, f x) = ∏ x in s, a ^ f x := +begin + induction s using finset.cons_induction with i s hi ihs, + { rw [sum_empty, finset.prod_empty, rpow_zero] }, + { rw forall_mem_cons at h, + rw [sum_cons, prod_cons, ← ihs h.2, rpow_add_of_nonneg ha h.1 (sum_nonneg h.2)] } +end + +lemma rpow_neg {x : ℝ} (hx : 0 ≤ x) (y : ℝ) : x ^ -y = (x ^ y)⁻¹ := +by simp only [rpow_def_of_nonneg hx]; split_ifs; simp [*, exp_neg] at * + +lemma rpow_sub {x : ℝ} (hx : 0 < x) (y z : ℝ) : x ^ (y - z) = x ^ y / x ^ z := +by simp only [sub_eq_add_neg, rpow_add hx, rpow_neg (le_of_lt hx), div_eq_mul_inv] + +lemma rpow_sub' {x : ℝ} (hx : 0 ≤ x) {y z : ℝ} (h : y - z ≠ 0) : + x ^ (y - z) = x ^ y / x ^ z := +by { simp only [sub_eq_add_neg] at h ⊢, simp only [rpow_add' hx h, rpow_neg hx, div_eq_mul_inv] } + +end real + +/-! +## Comparing real and complex powers +-/ + +namespace complex + +lemma of_real_cpow {x : ℝ} (hx : 0 ≤ x) (y : ℝ) : ((x ^ y : ℝ) : ℂ) = (x : ℂ) ^ (y : ℂ) := +by simp only [real.rpow_def_of_nonneg hx, complex.cpow_def, of_real_eq_zero]; split_ifs; + simp [complex.of_real_log hx] + +lemma of_real_cpow_of_nonpos {x : ℝ} (hx : x ≤ 0) (y : ℂ) : + (x : ℂ) ^ y = ((-x) : ℂ) ^ y * exp (π * I * y) := +begin + rcases hx.eq_or_lt with rfl|hlt, + { rcases eq_or_ne y 0 with rfl|hy; simp * }, + have hne : (x : ℂ) ≠ 0, from of_real_ne_zero.mpr hlt.ne, + rw [cpow_def_of_ne_zero hne, cpow_def_of_ne_zero (neg_ne_zero.2 hne), ← exp_add, ← add_mul, + log, log, abs.map_neg, arg_of_real_of_neg hlt, ← of_real_neg, + arg_of_real_of_nonneg (neg_nonneg.2 hx), of_real_zero, zero_mul, add_zero] +end + +lemma abs_cpow_of_ne_zero {z : ℂ} (hz : z ≠ 0) (w : ℂ) : + abs (z ^ w) = abs z ^ w.re / real.exp (arg z * im w) := +by rw [cpow_def_of_ne_zero hz, abs_exp, mul_re, log_re, log_im, real.exp_sub, + real.rpow_def_of_pos (abs.pos hz)] + +lemma abs_cpow_of_imp {z w : ℂ} (h : z = 0 → w.re = 0 → w = 0) : + abs (z ^ w) = abs z ^ w.re / real.exp (arg z * im w) := +begin + rcases ne_or_eq z 0 with hz|rfl; [exact (abs_cpow_of_ne_zero hz w), rw map_zero], + cases eq_or_ne w.re 0 with hw hw, + { simp [hw, h rfl hw] }, + { rw [real.zero_rpow hw, zero_div, zero_cpow, map_zero], + exact ne_of_apply_ne re hw } +end + +lemma abs_cpow_le (z w : ℂ) : abs (z ^ w) ≤ abs z ^ w.re / real.exp (arg z * im w) := +begin + rcases ne_or_eq z 0 with hz|rfl; [exact (abs_cpow_of_ne_zero hz w).le, rw map_zero], + rcases eq_or_ne w 0 with rfl|hw, { simp }, + rw [zero_cpow hw, map_zero], + exact div_nonneg (real.rpow_nonneg_of_nonneg le_rfl _) (real.exp_pos _).le +end + +@[simp] lemma abs_cpow_real (x : ℂ) (y : ℝ) : abs (x ^ (y : ℂ)) = x.abs ^ y := +by rcases eq_or_ne x 0 with rfl|hx; [rcases eq_or_ne y 0 with rfl|hy, skip]; + simp [*, abs_cpow_of_ne_zero] + +@[simp] lemma abs_cpow_inv_nat (x : ℂ) (n : ℕ) : abs (x ^ (n⁻¹ : ℂ)) = x.abs ^ (n⁻¹ : ℝ) := +by rw ← abs_cpow_real; simp [-abs_cpow_real] + +lemma abs_cpow_eq_rpow_re_of_pos {x : ℝ} (hx : 0 < x) (y : ℂ) : abs (x ^ y) = x ^ y.re := +by rw [abs_cpow_of_ne_zero (of_real_ne_zero.mpr hx.ne'), arg_of_real_of_nonneg hx.le, zero_mul, + real.exp_zero, div_one, abs_of_nonneg hx.le] + +lemma abs_cpow_eq_rpow_re_of_nonneg {x : ℝ} (hx : 0 ≤ x) {y : ℂ} (hy : re y ≠ 0) : + abs (x ^ y) = x ^ re y := +begin + rcases hx.eq_or_lt with rfl|hlt, + { rw [of_real_zero, zero_cpow, map_zero, real.zero_rpow hy], + exact ne_of_apply_ne re hy }, + { exact abs_cpow_eq_rpow_re_of_pos hlt y } +end + +end complex + +/-! +## Further algebraic properties of `rpow` +-/ + +namespace real + +variables {x y z : ℝ} + +lemma rpow_mul {x : ℝ} (hx : 0 ≤ x) (y z : ℝ) : x ^ (y * z) = (x ^ y) ^ z := +by rw [← complex.of_real_inj, complex.of_real_cpow (rpow_nonneg_of_nonneg hx _), + complex.of_real_cpow hx, complex.of_real_mul, complex.cpow_mul, complex.of_real_cpow hx]; + simp only [(complex.of_real_mul _ _).symm, (complex.of_real_log hx).symm, + complex.of_real_im, neg_lt_zero, pi_pos, le_of_lt pi_pos] + +lemma rpow_add_int {x : ℝ} (hx : x ≠ 0) (y : ℝ) (n : ℤ) : x ^ (y + n) = x ^ y * x ^ n := +by rw [rpow_def, complex.of_real_add, complex.cpow_add _ _ (complex.of_real_ne_zero.mpr hx), + complex.of_real_int_cast, complex.cpow_int_cast, ← complex.of_real_zpow, mul_comm, + complex.of_real_mul_re, ← rpow_def, mul_comm] + +lemma rpow_add_nat {x : ℝ} (hx : x ≠ 0) (y : ℝ) (n : ℕ) : x ^ (y + n) = x ^ y * x ^ n := +by simpa using rpow_add_int hx y n + +lemma rpow_sub_int {x : ℝ} (hx : x ≠ 0) (y : ℝ) (n : ℤ) : x ^ (y - n) = x ^ y / x ^ n := +by simpa using rpow_add_int hx y (-n) + +lemma rpow_sub_nat {x : ℝ} (hx : x ≠ 0) (y : ℝ) (n : ℕ) : x ^ (y - n) = x ^ y / x ^ n := +by simpa using rpow_sub_int hx y n + +lemma rpow_add_one {x : ℝ} (hx : x ≠ 0) (y : ℝ) : x ^ (y + 1) = x ^ y * x := +by simpa using rpow_add_nat hx y 1 + +lemma rpow_sub_one {x : ℝ} (hx : x ≠ 0) (y : ℝ) : x ^ (y - 1) = x ^ y / x := +by simpa using rpow_sub_nat hx y 1 + +@[simp, norm_cast] lemma rpow_int_cast (x : ℝ) (n : ℤ) : x ^ (n : ℝ) = x ^ n := +by simp only [rpow_def, ← complex.of_real_zpow, complex.cpow_int_cast, + complex.of_real_int_cast, complex.of_real_re] + +@[simp, norm_cast] lemma rpow_nat_cast (x : ℝ) (n : ℕ) : x ^ (n : ℝ) = x ^ n := +by simpa using rpow_int_cast x n + +@[simp] lemma rpow_two (x : ℝ) : x ^ (2 : ℝ) = x ^ 2 := +by { rw ← rpow_nat_cast, simp only [nat.cast_bit0, nat.cast_one] } + +lemma rpow_neg_one (x : ℝ) : x ^ (-1 : ℝ) = x⁻¹ := +begin + suffices H : x ^ ((-1 : ℤ) : ℝ) = x⁻¹, by rwa [int.cast_neg, int.cast_one] at H, + simp only [rpow_int_cast, zpow_one, zpow_neg], +end + +lemma mul_rpow {x y z : ℝ} (h : 0 ≤ x) (h₁ : 0 ≤ y) : (x*y)^z = x^z * y^z := +begin + iterate 3 { rw real.rpow_def_of_nonneg }, split_ifs; simp * at *, + { have hx : 0 < x, + { cases lt_or_eq_of_le h with h₂ h₂, { exact h₂ }, + exfalso, apply h_2, exact eq.symm h₂ }, + have hy : 0 < y, + { cases lt_or_eq_of_le h₁ with h₂ h₂, { exact h₂ }, + exfalso, apply h_3, exact eq.symm h₂ }, + rw [log_mul (ne_of_gt hx) (ne_of_gt hy), add_mul, exp_add]}, + { exact h₁ }, + { exact h }, + { exact mul_nonneg h h₁ }, +end + +lemma inv_rpow (hx : 0 ≤ x) (y : ℝ) : (x⁻¹)^y = (x^y)⁻¹ := +by simp only [← rpow_neg_one, ← rpow_mul hx, mul_comm] + +lemma div_rpow (hx : 0 ≤ x) (hy : 0 ≤ y) (z : ℝ) : (x / y) ^ z = x^z / y^z := +by simp only [div_eq_mul_inv, mul_rpow hx (inv_nonneg.2 hy), inv_rpow hy] + +lemma log_rpow {x : ℝ} (hx : 0 < x) (y : ℝ) : log (x^y) = y * (log x) := +begin + apply exp_injective, + rw [exp_log (rpow_pos_of_pos hx y), ← exp_log hx, mul_comm, rpow_def_of_pos (exp_pos (log x)) y], +end + +/-! +## Order and monotonicity +-/ + +lemma rpow_lt_rpow (hx : 0 ≤ x) (hxy : x < y) (hz : 0 < z) : x^z < y^z := +begin + rw le_iff_eq_or_lt at hx, cases hx, + { rw [← hx, zero_rpow (ne_of_gt hz)], exact rpow_pos_of_pos (by rwa ← hx at hxy) _ }, + rw [rpow_def_of_pos hx, rpow_def_of_pos (lt_trans hx hxy), exp_lt_exp], + exact mul_lt_mul_of_pos_right (log_lt_log hx hxy) hz +end + +lemma rpow_le_rpow {x y z: ℝ} (h : 0 ≤ x) (h₁ : x ≤ y) (h₂ : 0 ≤ z) : x^z ≤ y^z := +begin + rcases eq_or_lt_of_le h₁ with rfl|h₁', { refl }, + rcases eq_or_lt_of_le h₂ with rfl|h₂', { simp }, + exact le_of_lt (rpow_lt_rpow h h₁' h₂') +end + +lemma rpow_lt_rpow_iff (hx : 0 ≤ x) (hy : 0 ≤ y) (hz : 0 < z) : x ^ z < y ^ z ↔ x < y := +⟨lt_imp_lt_of_le_imp_le $ λ h, rpow_le_rpow hy h (le_of_lt hz), λ h, rpow_lt_rpow hx h hz⟩ + +lemma rpow_le_rpow_iff (hx : 0 ≤ x) (hy : 0 ≤ y) (hz : 0 < z) : x ^ z ≤ y ^ z ↔ x ≤ y := +le_iff_le_iff_lt_iff_lt.2 $ rpow_lt_rpow_iff hy hx hz + +lemma le_rpow_inv_iff_of_neg (hx : 0 < x) (hy : 0 < y) (hz : z < 0) : + x ≤ y ^ z⁻¹ ↔ y ≤ x ^ z := +begin + have hz' : 0 < -z := by rwa [lt_neg, neg_zero], + have hxz : 0 < x ^ (-z) := real.rpow_pos_of_pos hx _, + have hyz : 0 < y ^ z⁻¹ := real.rpow_pos_of_pos hy _, + rw [←real.rpow_le_rpow_iff hx.le hyz.le hz', ←real.rpow_mul hy.le], + simp only [ne_of_lt hz, real.rpow_neg_one, mul_neg, inv_mul_cancel, ne.def, not_false_iff], + rw [le_inv hxz hy, ←real.rpow_neg_one, ←real.rpow_mul hx.le], + simp, +end + +lemma lt_rpow_inv_iff_of_neg (hx : 0 < x) (hy : 0 < y) (hz : z < 0) : + x < y ^ z⁻¹ ↔ y < x ^ z := +begin + have hz' : 0 < -z := by rwa [lt_neg, neg_zero], + have hxz : 0 < x ^ (-z) := real.rpow_pos_of_pos hx _, + have hyz : 0 < y ^ z⁻¹ := real.rpow_pos_of_pos hy _, + rw [←real.rpow_lt_rpow_iff hx.le hyz.le hz', ←real.rpow_mul hy.le], + simp only [ne_of_lt hz, real.rpow_neg_one, mul_neg, inv_mul_cancel, ne.def, not_false_iff], + rw [lt_inv hxz hy, ←real.rpow_neg_one, ←real.rpow_mul hx.le], + simp, +end + +lemma rpow_inv_lt_iff_of_neg (hx : 0 < x) (hy : 0 < y) (hz : z < 0) : + x ^ z⁻¹ < y ↔ y ^ z < x := +begin + convert lt_rpow_inv_iff_of_neg (real.rpow_pos_of_pos hx _) (real.rpow_pos_of_pos hy _) hz; + simp [←real.rpow_mul hx.le, ←real.rpow_mul hy.le, ne_of_lt hz], +end + +lemma rpow_inv_le_iff_of_neg (hx : 0 < x) (hy : 0 < y) (hz : z < 0) : + x ^ z⁻¹ ≤ y ↔ y ^ z ≤ x := +begin + convert le_rpow_inv_iff_of_neg (real.rpow_pos_of_pos hx _) (real.rpow_pos_of_pos hy _) hz; + simp [←real.rpow_mul hx.le, ←real.rpow_mul hy.le, ne_of_lt hz], +end + +lemma rpow_lt_rpow_of_exponent_lt (hx : 1 < x) (hyz : y < z) : x^y < x^z := +begin + repeat {rw [rpow_def_of_pos (lt_trans zero_lt_one hx)]}, + rw exp_lt_exp, exact mul_lt_mul_of_pos_left hyz (log_pos hx), +end + +lemma rpow_le_rpow_of_exponent_le (hx : 1 ≤ x) (hyz : y ≤ z) : x^y ≤ x^z := +begin + repeat {rw [rpow_def_of_pos (lt_of_lt_of_le zero_lt_one hx)]}, + rw exp_le_exp, exact mul_le_mul_of_nonneg_left hyz (log_nonneg hx), +end + +@[simp] lemma rpow_le_rpow_left_iff (hx : 1 < x) : x ^ y ≤ x ^ z ↔ y ≤ z := +begin + have x_pos : 0 < x := lt_trans zero_lt_one hx, + rw [←log_le_log (rpow_pos_of_pos x_pos y) (rpow_pos_of_pos x_pos z), + log_rpow x_pos, log_rpow x_pos, mul_le_mul_right (log_pos hx)], +end + +@[simp] lemma rpow_lt_rpow_left_iff (hx : 1 < x) : x ^ y < x ^ z ↔ y < z := +by rw [lt_iff_not_le, rpow_le_rpow_left_iff hx, lt_iff_not_le] + +lemma rpow_lt_rpow_of_exponent_gt (hx0 : 0 < x) (hx1 : x < 1) (hyz : z < y) : + x^y < x^z := +begin + repeat {rw [rpow_def_of_pos hx0]}, + rw exp_lt_exp, exact mul_lt_mul_of_neg_left hyz (log_neg hx0 hx1), +end + +lemma rpow_le_rpow_of_exponent_ge (hx0 : 0 < x) (hx1 : x ≤ 1) (hyz : z ≤ y) : + x^y ≤ x^z := +begin + repeat {rw [rpow_def_of_pos hx0]}, + rw exp_le_exp, exact mul_le_mul_of_nonpos_left hyz (log_nonpos (le_of_lt hx0) hx1), +end + +@[simp] lemma rpow_le_rpow_left_iff_of_base_lt_one (hx0 : 0 < x) (hx1 : x < 1) : + x ^ y ≤ x ^ z ↔ z ≤ y := +begin + rw [←log_le_log (rpow_pos_of_pos hx0 y) (rpow_pos_of_pos hx0 z), + log_rpow hx0, log_rpow hx0, mul_le_mul_right_of_neg (log_neg hx0 hx1)], +end + +@[simp] lemma rpow_lt_rpow_left_iff_of_base_lt_one (hx0 : 0 < x) (hx1 : x < 1) : + x ^ y < x ^ z ↔ z < y := +by rw [lt_iff_not_le, rpow_le_rpow_left_iff_of_base_lt_one hx0 hx1, lt_iff_not_le] + +lemma rpow_lt_one {x z : ℝ} (hx1 : 0 ≤ x) (hx2 : x < 1) (hz : 0 < z) : x^z < 1 := +by { rw ← one_rpow z, exact rpow_lt_rpow hx1 hx2 hz } + +lemma rpow_le_one {x z : ℝ} (hx1 : 0 ≤ x) (hx2 : x ≤ 1) (hz : 0 ≤ z) : x^z ≤ 1 := +by { rw ← one_rpow z, exact rpow_le_rpow hx1 hx2 hz } + +lemma rpow_lt_one_of_one_lt_of_neg {x z : ℝ} (hx : 1 < x) (hz : z < 0) : x^z < 1 := +by { convert rpow_lt_rpow_of_exponent_lt hx hz, exact (rpow_zero x).symm } + +lemma rpow_le_one_of_one_le_of_nonpos {x z : ℝ} (hx : 1 ≤ x) (hz : z ≤ 0) : x^z ≤ 1 := +by { convert rpow_le_rpow_of_exponent_le hx hz, exact (rpow_zero x).symm } + +lemma one_lt_rpow {x z : ℝ} (hx : 1 < x) (hz : 0 < z) : 1 < x^z := +by { rw ← one_rpow z, exact rpow_lt_rpow zero_le_one hx hz } + +lemma one_le_rpow {x z : ℝ} (hx : 1 ≤ x) (hz : 0 ≤ z) : 1 ≤ x^z := +by { rw ← one_rpow z, exact rpow_le_rpow zero_le_one hx hz } + +lemma one_lt_rpow_of_pos_of_lt_one_of_neg (hx1 : 0 < x) (hx2 : x < 1) (hz : z < 0) : + 1 < x^z := +by { convert rpow_lt_rpow_of_exponent_gt hx1 hx2 hz, exact (rpow_zero x).symm } + +lemma one_le_rpow_of_pos_of_le_one_of_nonpos (hx1 : 0 < x) (hx2 : x ≤ 1) (hz : z ≤ 0) : + 1 ≤ x^z := +by { convert rpow_le_rpow_of_exponent_ge hx1 hx2 hz, exact (rpow_zero x).symm } + +lemma rpow_lt_one_iff_of_pos (hx : 0 < x) : x ^ y < 1 ↔ 1 < x ∧ y < 0 ∨ x < 1 ∧ 0 < y := +by rw [rpow_def_of_pos hx, exp_lt_one_iff, mul_neg_iff, log_pos_iff hx, log_neg_iff hx] + +lemma rpow_lt_one_iff (hx : 0 ≤ x) : x ^ y < 1 ↔ x = 0 ∧ y ≠ 0 ∨ 1 < x ∧ y < 0 ∨ x < 1 ∧ 0 < y := +begin + rcases hx.eq_or_lt with (rfl|hx), + { rcases em (y = 0) with (rfl|hy); simp [*, lt_irrefl, zero_lt_one] }, + { simp [rpow_lt_one_iff_of_pos hx, hx.ne.symm] } +end + +lemma one_lt_rpow_iff_of_pos (hx : 0 < x) : 1 < x ^ y ↔ 1 < x ∧ 0 < y ∨ x < 1 ∧ y < 0 := +by rw [rpow_def_of_pos hx, one_lt_exp_iff, mul_pos_iff, log_pos_iff hx, log_neg_iff hx] + +lemma one_lt_rpow_iff (hx : 0 ≤ x) : 1 < x ^ y ↔ 1 < x ∧ 0 < y ∨ 0 < x ∧ x < 1 ∧ y < 0 := +begin + rcases hx.eq_or_lt with (rfl|hx), + { rcases em (y = 0) with (rfl|hy); simp [*, lt_irrefl, (zero_lt_one' ℝ).not_lt] }, + { simp [one_lt_rpow_iff_of_pos hx, hx] } +end + +lemma rpow_le_rpow_of_exponent_ge' (hx0 : 0 ≤ x) (hx1 : x ≤ 1) (hz : 0 ≤ z) (hyz : z ≤ y) : + x^y ≤ x^z := +begin + rcases eq_or_lt_of_le hx0 with rfl | hx0', + { rcases eq_or_lt_of_le hz with rfl | hz', + { exact (rpow_zero 0).symm ▸ (rpow_le_one hx0 hx1 hyz), }, + rw [zero_rpow, zero_rpow]; linarith, }, + { exact rpow_le_rpow_of_exponent_ge hx0' hx1 hyz, }, +end + +lemma rpow_left_inj_on {x : ℝ} (hx : x ≠ 0) : + inj_on (λ y : ℝ, y^x) {y : ℝ | 0 ≤ y} := +begin + rintros y hy z hz (hyz : y ^ x = z ^ x), + rw [←rpow_one y, ←rpow_one z, ←_root_.mul_inv_cancel hx, rpow_mul hy, rpow_mul hz, hyz] +end + +lemma le_rpow_iff_log_le (hx : 0 < x) (hy : 0 < y) : + x ≤ y^z ↔ real.log x ≤ z * real.log y := +by rw [←real.log_le_log hx (real.rpow_pos_of_pos hy z), real.log_rpow hy] + +lemma le_rpow_of_log_le (hx : 0 ≤ x) (hy : 0 < y) (h : real.log x ≤ z * real.log y) : + x ≤ y^z := +begin + obtain hx | rfl := hx.lt_or_eq, + { exact (le_rpow_iff_log_le hx hy).2 h }, + exact (real.rpow_pos_of_pos hy z).le, +end + +lemma lt_rpow_iff_log_lt (hx : 0 < x) (hy : 0 < y) : + x < y^z ↔ real.log x < z * real.log y := +by rw [←real.log_lt_log_iff hx (real.rpow_pos_of_pos hy z), real.log_rpow hy] + +lemma lt_rpow_of_log_lt (hx : 0 ≤ x) (hy : 0 < y) (h : real.log x < z * real.log y) : + x < y^z := +begin + obtain hx | rfl := hx.lt_or_eq, + { exact (lt_rpow_iff_log_lt hx hy).2 h }, + exact real.rpow_pos_of_pos hy z, +end + +lemma rpow_le_one_iff_of_pos (hx : 0 < x) : x ^ y ≤ 1 ↔ 1 ≤ x ∧ y ≤ 0 ∨ x ≤ 1 ∧ 0 ≤ y := +by rw [rpow_def_of_pos hx, exp_le_one_iff, mul_nonpos_iff, log_nonneg_iff hx, log_nonpos_iff hx] + +/-- Bound for `|log x * x ^ t|` in the interval `(0, 1]`, for positive real `t`. -/ +lemma abs_log_mul_self_rpow_lt (x t : ℝ) (h1 : 0 < x) (h2 : x ≤ 1) (ht : 0 < t) : + |log x * x ^ t| < 1 / t := +begin + rw lt_div_iff ht, + have := abs_log_mul_self_lt (x ^ t) (rpow_pos_of_pos h1 t) (rpow_le_one h1.le h2 ht.le), + rwa [log_rpow h1, mul_assoc, abs_mul, abs_of_pos ht, mul_comm] at this +end + +lemma pow_nat_rpow_nat_inv {x : ℝ} (hx : 0 ≤ x) {n : ℕ} (hn : n ≠ 0) : + (x ^ n) ^ (n⁻¹ : ℝ) = x := +have hn0 : (n : ℝ) ≠ 0, from nat.cast_ne_zero.2 hn, +by rw [← rpow_nat_cast, ← rpow_mul hx, mul_inv_cancel hn0, rpow_one] + +lemma rpow_nat_inv_pow_nat {x : ℝ} (hx : 0 ≤ x) {n : ℕ} (hn : n ≠ 0) : + (x ^ (n⁻¹ : ℝ)) ^ n = x := +have hn0 : (n : ℝ) ≠ 0, from nat.cast_ne_zero.2 hn, +by rw [← rpow_nat_cast, ← rpow_mul hx, inv_mul_cancel hn0, rpow_one] + + +end real + +/-! +## Square roots of reals +-/ +namespace real + +variables {z x y : ℝ} + +section sqrt + +lemma sqrt_eq_rpow (x : ℝ) : sqrt x = x ^ (1/(2:ℝ)) := +begin + obtain h | h := le_or_lt 0 x, + { rw [← mul_self_inj_of_nonneg (sqrt_nonneg _) (rpow_nonneg_of_nonneg h _), mul_self_sqrt h, + ← sq, ← rpow_nat_cast, ← rpow_mul h], + norm_num }, + { have : 1 / (2:ℝ) * π = π / (2:ℝ), ring, + rw [sqrt_eq_zero_of_nonpos h.le, rpow_def_of_neg h, this, cos_pi_div_two, mul_zero] } +end + +lemma rpow_div_two_eq_sqrt {x : ℝ} (r : ℝ) (hx : 0 ≤ x) : x ^ (r/2) = (sqrt x) ^ r := +begin + rw [sqrt_eq_rpow, ← rpow_mul hx], + congr, + ring, +end + +end sqrt + +variables {n : ℕ} + +lemma exists_rat_pow_btwn_rat_aux (hn : n ≠ 0) (x y : ℝ) (h : x < y) (hy : 0 < y) : + ∃ q : ℚ, 0 < q ∧ x < q^n ∧ ↑q^n < y := +begin + have hn' : 0 < (n : ℝ) := by exact_mod_cast hn.bot_lt, + obtain ⟨q, hxq, hqy⟩ := exists_rat_btwn (rpow_lt_rpow (le_max_left 0 x) (max_lt hy h) $ + inv_pos.mpr hn'), + have := rpow_nonneg_of_nonneg (le_max_left 0 x) n⁻¹, + have hq := this.trans_lt hxq, + replace hxq := rpow_lt_rpow this hxq hn', + replace hqy := rpow_lt_rpow hq.le hqy hn', + rw [rpow_nat_cast, rpow_nat_cast, rpow_nat_inv_pow_nat _ hn] at hxq hqy, + exact ⟨q, by exact_mod_cast hq, (le_max_right _ _).trans_lt hxq, hqy⟩, + { exact le_max_left _ _ }, + { exact hy.le } +end + +lemma exists_rat_pow_btwn_rat (hn : n ≠ 0) {x y : ℚ} (h : x < y) (hy : 0 < y) : + ∃ q : ℚ, 0 < q ∧ x < q^n ∧ q^n < y := +by apply_mod_cast exists_rat_pow_btwn_rat_aux hn x y; assumption + +/-- There is a rational power between any two positive elements of an archimedean ordered field. -/ +lemma exists_rat_pow_btwn {α : Type*} [linear_ordered_field α] [archimedean α] (hn : n ≠ 0) + {x y : α} (h : x < y) (hy : 0 < y) : ∃ q : ℚ, 0 < q ∧ x < q^n ∧ (q^n : α) < y := +begin + obtain ⟨q₂, hx₂, hy₂⟩ := exists_rat_btwn (max_lt h hy), + obtain ⟨q₁, hx₁, hq₁₂⟩ := exists_rat_btwn hx₂, + have : (0 : α) < q₂ := (le_max_right _ _).trans_lt hx₂, + norm_cast at hq₁₂ this, + obtain ⟨q, hq, hq₁, hq₂⟩ := exists_rat_pow_btwn_rat hn hq₁₂ this, + refine ⟨q, hq, (le_max_left _ _).trans_lt $ hx₁.trans _, hy₂.trans' _⟩; assumption_mod_cast, +end + +end real diff --git a/src/analysis/special_functions/pow/tactic.lean b/src/analysis/special_functions/pow/tactic.lean new file mode 100644 index 0000000000000..1a426e55b07c8 --- /dev/null +++ b/src/analysis/special_functions/pow/tactic.lean @@ -0,0 +1,213 @@ +/- +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, Sébastien Gouëzel, + Rémy Degenne, David Loeffler +-/ +import analysis.special_functions.pow.nnreal + +/-! +# Tactics for power functions + +This file contains extensions to the `norm_num` and `positivity` tactics to handle power operations +on `ℂ`, `ℝ`, `ℝ≥0`, and `ℝ≥0∞`. + +TODO: Split up the contents of this file and merge with other files in +`analysis/special_functions/pow/`, to keep the tactics together with the corresponding definitions. + +-/ + +open_locale nnreal ennreal + +/-! +## Complex case +-/ + +namespace norm_num + +theorem cpow_pos (a b : ℂ) (b' : ℕ) (c : ℂ) (hb : b = b') (h : a ^ b' = c) : a ^ b = c := +by rw [← h, hb, complex.cpow_nat_cast] + +theorem cpow_neg (a b : ℂ) (b' : ℕ) (c c' : ℂ) + (hb : b = b') (h : a ^ b' = c) (hc : c⁻¹ = c') : a ^ -b = c' := +by rw [← hc, ← h, hb, complex.cpow_neg, complex.cpow_nat_cast] + +end norm_num + +/-! +## Real case +-/ + +namespace norm_num +open tactic + +theorem rpow_pos (a b : ℝ) (b' : ℕ) (c : ℝ) (hb : (b':ℝ) = b) (h : a ^ b' = c) : a ^ b = c := +by rw [← h, ← hb, real.rpow_nat_cast] + +theorem rpow_neg (a b : ℝ) (b' : ℕ) (c c' : ℝ) + (a0 : 0 ≤ a) (hb : (b':ℝ) = b) (h : a ^ b' = c) (hc : c⁻¹ = c') : a ^ -b = c' := +by rw [← hc, ← h, ← hb, real.rpow_neg a0, real.rpow_nat_cast] + +/-- Evaluate `real.rpow a b` where `a` is a rational numeral and `b` is an integer. +(This cannot go via the generalized version `prove_rpow'` because `rpow_pos` has a side condition; +we do not attempt to evaluate `a ^ b` where `a` and `b` are both negative because it comes +out to some garbage.) -/ +meta def prove_rpow (a b : expr) : tactic (expr × expr) := do + na ← a.to_rat, + ic ← mk_instance_cache `(ℝ), + match match_sign b with + | sum.inl b := do + (ic, a0) ← guard (na ≥ 0) >> prove_nonneg ic a, + nc ← mk_instance_cache `(ℕ), + (ic, nc, b', hb) ← prove_nat_uncast ic nc b, + (ic, c, h) ← prove_pow a na ic b', + cr ← c.to_rat, + (ic, c', hc) ← prove_inv ic c cr, + pure (c', (expr.const ``rpow_neg []).mk_app [a, b, b', c, c', a0, hb, h, hc]) + | sum.inr ff := pure (`(1:ℝ), expr.const ``real.rpow_zero [] a) + | sum.inr tt := do + nc ← mk_instance_cache `(ℕ), + (ic, nc, b', hb) ← prove_nat_uncast ic nc b, + (ic, c, h) ← prove_pow a na ic b', + pure (c, (expr.const ``rpow_pos []).mk_app [a, b, b', c, hb, h]) + end + +end norm_num + +namespace tactic +namespace positivity + +/-- Auxiliary definition for the `positivity` tactic to handle real powers of reals. -/ +meta def prove_rpow (a b : expr) : tactic strictness := +do + strictness_a ← core a, + match strictness_a with + | nonnegative p := nonnegative <$> mk_app ``real.rpow_nonneg_of_nonneg [p, b] + | positive p := positive <$> mk_app ``real.rpow_pos_of_pos [p, b] + | _ := failed + end + +end positivity +end tactic + +/-! +## General-purpose tactics + +The following code covers all the cases `ℂ`, `ℝ`, `ℝ≥0`, and `ℝ≥0∞` together. +-/ + +namespace norm_num +open tactic + +/-- Generalized version of `prove_cpow`, `prove_nnrpow`, `prove_ennrpow`. -/ +meta def prove_rpow' (pos neg zero : name) (α β one a b : expr) : tactic (expr × expr) := do + na ← a.to_rat, + icα ← mk_instance_cache α, + icβ ← mk_instance_cache β, + match match_sign b with + | sum.inl b := do + nc ← mk_instance_cache `(ℕ), + (icβ, nc, b', hb) ← prove_nat_uncast icβ nc b, + (icα, c, h) ← prove_pow a na icα b', + cr ← c.to_rat, + (icα, c', hc) ← prove_inv icα c cr, + pure (c', (expr.const neg []).mk_app [a, b, b', c, c', hb, h, hc]) + | sum.inr ff := pure (one, expr.const zero [] a) + | sum.inr tt := do + nc ← mk_instance_cache `(ℕ), + (icβ, nc, b', hb) ← prove_nat_uncast icβ nc b, + (icα, c, h) ← prove_pow a na icα b', + pure (c, (expr.const pos []).mk_app [a, b, b', c, hb, h]) + end + +open_locale nnreal ennreal + +theorem nnrpow_pos (a : ℝ≥0) (b : ℝ) (b' : ℕ) (c : ℝ≥0) + (hb : b = b') (h : a ^ b' = c) : a ^ b = c := +by rw [← h, hb, nnreal.rpow_nat_cast] + +theorem nnrpow_neg (a : ℝ≥0) (b : ℝ) (b' : ℕ) (c c' : ℝ≥0) + (hb : b = b') (h : a ^ b' = c) (hc : c⁻¹ = c') : a ^ -b = c' := +by rw [← hc, ← h, hb, nnreal.rpow_neg, nnreal.rpow_nat_cast] + +theorem ennrpow_pos (a : ℝ≥0∞) (b : ℝ) (b' : ℕ) (c : ℝ≥0∞) + (hb : b = b') (h : a ^ b' = c) : a ^ b = c := +by rw [← h, hb, ennreal.rpow_nat_cast] + +theorem ennrpow_neg (a : ℝ≥0∞) (b : ℝ) (b' : ℕ) (c c' : ℝ≥0∞) + (hb : b = b') (h : a ^ b' = c) (hc : c⁻¹ = c') : a ^ -b = c' := +by rw [← hc, ← h, hb, ennreal.rpow_neg, ennreal.rpow_nat_cast] + +/-- Evaluate `complex.cpow a b` where `a` is a rational numeral and `b` is an integer. -/ +meta def prove_cpow : expr → expr → tactic (expr × expr) := +prove_rpow' ``cpow_pos ``cpow_neg ``complex.cpow_zero `(ℂ) `(ℂ) `(1:ℂ) + +/-- Evaluate `nnreal.rpow a b` where `a` is a rational numeral and `b` is an integer. -/ +meta def prove_nnrpow : expr → expr → tactic (expr × expr) := +prove_rpow' ``nnrpow_pos ``nnrpow_neg ``nnreal.rpow_zero `(ℝ≥0) `(ℝ) `(1:ℝ≥0) + +/-- Evaluate `ennreal.rpow a b` where `a` is a rational numeral and `b` is an integer. -/ +meta def prove_ennrpow : expr → expr → tactic (expr × expr) := +prove_rpow' ``ennrpow_pos ``ennrpow_neg ``ennreal.rpow_zero `(ℝ≥0∞) `(ℝ) `(1:ℝ≥0∞) + +/-- Evaluates expressions of the form `rpow a b`, `cpow a b` and `a ^ b` in the special case where +`b` is an integer and `a` is a positive rational (so it's really just a rational power). -/ +@[norm_num] meta def eval_rpow_cpow : expr → tactic (expr × expr) +| `(@has_pow.pow _ _ real.has_pow %%a %%b) := b.to_int >> prove_rpow a b +| `(real.rpow %%a %%b) := b.to_int >> prove_rpow a b +| `(@has_pow.pow _ _ complex.has_pow %%a %%b) := b.to_int >> prove_cpow a b +| `(complex.cpow %%a %%b) := b.to_int >> prove_cpow a b +| `(@has_pow.pow _ _ nnreal.real.has_pow %%a %%b) := b.to_int >> prove_nnrpow a b +| `(nnreal.rpow %%a %%b) := b.to_int >> prove_nnrpow a b +| `(@has_pow.pow _ _ ennreal.real.has_pow %%a %%b) := b.to_int >> prove_ennrpow a b +| `(ennreal.rpow %%a %%b) := b.to_int >> prove_ennrpow a b +| _ := tactic.failed + +end norm_num + +namespace tactic +namespace positivity + +private lemma nnrpow_pos {a : ℝ≥0} (ha : 0 < a) (b : ℝ) : 0 < a ^ b := nnreal.rpow_pos ha + +/-- Auxiliary definition for the `positivity` tactic to handle real powers of nonnegative reals. -/ +meta def prove_nnrpow (a b : expr) : tactic strictness := +do + strictness_a ← core a, + match strictness_a with + | positive p := positive <$> mk_app ``nnrpow_pos [p, b] + | _ := failed -- We already know `0 ≤ x` for all `x : ℝ≥0` + end + +private lemma ennrpow_pos {a : ℝ≥0∞} {b : ℝ} (ha : 0 < a) (hb : 0 < b) : 0 < a ^ b := +ennreal.rpow_pos_of_nonneg ha hb.le + +/-- Auxiliary definition for the `positivity` tactic to handle real powers of extended nonnegative +reals. -/ +meta def prove_ennrpow (a b : expr) : tactic strictness := +do + strictness_a ← core a, + strictness_b ← core b, + match strictness_a, strictness_b with + | positive pa, positive pb := positive <$> mk_app ``ennrpow_pos [pa, pb] + | positive pa, nonnegative pb := positive <$> mk_app ``ennreal.rpow_pos_of_nonneg [pa, pb] + | _, _ := failed -- We already know `0 ≤ x` for all `x : ℝ≥0∞` + end + +end positivity + +open positivity + +/-- Extension for the `positivity` tactic: exponentiation by a real number is nonnegative when the +base is nonnegative and positive when the base is positive. -/ +@[positivity] +meta def positivity_rpow : expr → tactic strictness +| `(@has_pow.pow _ _ real.has_pow %%a %%b) := prove_rpow a b +| `(real.rpow %%a %%b) := prove_rpow a b +| `(@has_pow.pow _ _ nnreal.real.has_pow %%a %%b) := prove_nnrpow a b +| `(nnreal.rpow %%a %%b) := prove_nnrpow a b +| `(@has_pow.pow _ _ ennreal.real.has_pow %%a %%b) := prove_ennrpow a b +| `(ennreal.rpow %%a %%b) := prove_ennrpow a b +| _ := failed + +end tactic diff --git a/src/analysis/specific_limits/floor_pow.lean b/src/analysis/specific_limits/floor_pow.lean index 9ef025f425f28..10db01964533b 100644 --- a/src/analysis/specific_limits/floor_pow.lean +++ b/src/analysis/specific_limits/floor_pow.lean @@ -5,7 +5,7 @@ Authors: Sébastien Gouëzel -/ import analysis.specific_limits.basic -import analysis.special_functions.pow +import analysis.special_functions.pow.real /-! # Results on discretized exponentials diff --git a/src/combinatorics/additive/behrend.lean b/src/combinatorics/additive/behrend.lean index 857a123b6c14e..6e4eae183ca42 100644 --- a/src/combinatorics/additive/behrend.lean +++ b/src/combinatorics/additive/behrend.lean @@ -7,6 +7,7 @@ import analysis.inner_product_space.pi_L2 import combinatorics.additive.salem_spencer import combinatorics.pigeonhole import data.complex.exponential_bounds +import analysis.special_functions.pow.tactic /-! # Behrend's bound on Roth numbers diff --git a/src/combinatorics/simple_graph/regularity/bound.lean b/src/combinatorics/simple_graph/regularity/bound.lean index 0142d46464d4b..756d5fb7b80ea 100644 --- a/src/combinatorics/simple_graph/regularity/bound.lean +++ b/src/combinatorics/simple_graph/regularity/bound.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies, Bhavik Mehta -/ import algebra.order.chebyshev -import analysis.special_functions.pow +import analysis.special_functions.pow.tactic import order.partition.equipartition /-! diff --git a/src/linear_algebra/quadratic_form/complex.lean b/src/linear_algebra/quadratic_form/complex.lean index 0a6dfd3393cd8..8d7c5c33434fb 100644 --- a/src/linear_algebra/quadratic_form/complex.lean +++ b/src/linear_algebra/quadratic_form/complex.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Anne Baanen, Kexing Ying, Eric Wieser -/ import linear_algebra.quadratic_form.isometry -import analysis.special_functions.pow +import analysis.special_functions.pow.complex /-! # Quadratic forms over the complex numbers diff --git a/src/linear_algebra/quadratic_form/real.lean b/src/linear_algebra/quadratic_form/real.lean index 8e172a1e9c1c5..289a2b2b2b552 100644 --- a/src/linear_algebra/quadratic_form/real.lean +++ b/src/linear_algebra/quadratic_form/real.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Anne Baanen, Kexing Ying, Eric Wieser -/ import linear_algebra.quadratic_form.isometry -import analysis.special_functions.pow +import analysis.special_functions.pow.real import data.real.sign /-! diff --git a/src/measure_theory/function/convergence_in_measure.lean b/src/measure_theory/function/convergence_in_measure.lean index 01d66d182cc2b..ee8737e30c14c 100644 --- a/src/measure_theory/function/convergence_in_measure.lean +++ b/src/measure_theory/function/convergence_in_measure.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Rémy Degenne, Kexing Ying. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne, Kexing Ying -/ -import analysis.special_functions.pow +import analysis.special_functions.pow.real import measure_theory.function.egorov import measure_theory.function.lp_space diff --git a/src/measure_theory/function/special_functions/basic.lean b/src/measure_theory/function/special_functions/basic.lean index 78d900f0c4bb9..66d515bbc5706 100644 --- a/src/measure_theory/function/special_functions/basic.lean +++ b/src/measure_theory/function/special_functions/basic.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ -import analysis.special_functions.pow +import analysis.special_functions.pow.nnreal import data.is_R_or_C.lemmas import measure_theory.constructions.borel_space.complex diff --git a/src/measure_theory/measure/hausdorff.lean b/src/measure_theory/measure/hausdorff.lean index 4c954fa1b605a..cf5ef8418662e 100644 --- a/src/measure_theory/measure/hausdorff.lean +++ b/src/measure_theory/measure/hausdorff.lean @@ -3,7 +3,6 @@ Copyright (c) 2021 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ -import analysis.special_functions.pow import measure_theory.constructions.borel_space.basic import measure_theory.measure.lebesgue import topology.metric_space.holder diff --git a/src/number_theory/bertrand.lean b/src/number_theory/bertrand.lean index b9dbb2af60d55..b98486b53e1d0 100644 --- a/src/number_theory/bertrand.lean +++ b/src/number_theory/bertrand.lean @@ -7,6 +7,7 @@ import data.nat.choose.factorization import data.nat.prime_norm_num import number_theory.primorial import analysis.convex.specific_functions +import analysis.special_functions.pow.tactic /-! # Bertrand's Postulate diff --git a/src/number_theory/class_number/admissible_card_pow_degree.lean b/src/number_theory/class_number/admissible_card_pow_degree.lean index 0566fcc9534bb..37a7b22bbc3d3 100644 --- a/src/number_theory/class_number/admissible_card_pow_degree.lean +++ b/src/number_theory/class_number/admissible_card_pow_degree.lean @@ -4,15 +4,15 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Anne Baanen -/ import number_theory.class_number.admissible_absolute_value -import analysis.special_functions.pow +import analysis.special_functions.pow.real import ring_theory.ideal.local_ring import data.polynomial.degree.card_pow_degree /-! # Admissible absolute values on polynomials -This file defines an admissible absolute value -`polynomial.card_pow_degree_is_admissible` which we use to show the class number -of the ring of integers of a function field is finite. + +This file defines an admissible absolute value `polynomial.card_pow_degree_is_admissible` which we +use to show the class number of the ring of integers of a function field is finite. ## Main results diff --git a/src/number_theory/class_number/finite.lean b/src/number_theory/class_number/finite.lean index 8406e7cdc4ba9..fb56d22dbdbb3 100644 --- a/src/number_theory/class_number/finite.lean +++ b/src/number_theory/class_number/finite.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Anne Baanen -/ -import analysis.special_functions.pow +import analysis.special_functions.pow.real import linear_algebra.free_module.pid import linear_algebra.matrix.absolute_value import number_theory.class_number.admissible_absolute_value diff --git a/src/number_theory/liouville/liouville_with.lean b/src/number_theory/liouville/liouville_with.lean index c05c8aeaa9979..a0f50ae4c916f 100644 --- a/src/number_theory/liouville/liouville_with.lean +++ b/src/number_theory/liouville/liouville_with.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Yury G. Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury G. Kudryashov -/ -import analysis.special_functions.pow +import analysis.special_functions.pow.asymptotics import number_theory.liouville.basic import topology.instances.irrational diff --git a/src/ring_theory/perfection.lean b/src/ring_theory/perfection.lean index 61ace5ec6eb22..8357e7a373fd7 100644 --- a/src/ring_theory/perfection.lean +++ b/src/ring_theory/perfection.lean @@ -8,7 +8,7 @@ import algebra.char_p.pi import algebra.char_p.quotient import algebra.char_p.subring import algebra.ring.pi -import analysis.special_functions.pow +import analysis.special_functions.pow.nnreal import field_theory.perfect_closure import ring_theory.localization.fraction_ring import ring_theory.subring.basic diff --git a/src/topology/metric_space/holder.lean b/src/topology/metric_space/holder.lean index 2db96800cd5eb..d7dbd7c8f7cb7 100644 --- a/src/topology/metric_space/holder.lean +++ b/src/topology/metric_space/holder.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury G. Kudryashov -/ import topology.metric_space.lipschitz -import analysis.special_functions.pow +import analysis.special_functions.pow.continuity /-! # Hölder continuous functions diff --git a/test/norm_num_ext.lean b/test/norm_num_ext.lean index d33578ed6e431..e1f93a9210b86 100644 --- a/test/norm_num_ext.lean +++ b/test/norm_num_ext.lean @@ -9,7 +9,7 @@ import data.int.gcd import data.nat.fib import data.nat.prime import data.nat.sqrt_norm_num -import analysis.special_functions.pow +import analysis.special_functions.pow.tactic import number_theory.legendre_symbol.norm_num /-! diff --git a/test/positivity.lean b/test/positivity.lean index fa2df257e3ffd..30b644d798dab 100644 --- a/test/positivity.lean +++ b/test/positivity.lean @@ -1,7 +1,7 @@ import algebra.order.interval import algebra.order.smul import analysis.normed.group.basic -import analysis.special_functions.pow +import analysis.special_functions.pow.tactic import combinatorics.simple_graph.density import data.complex.exponential import data.rat.nnrat From 4fa54b337f7d52805480306db1b1439c741848c8 Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Tue, 16 May 2023 00:23:25 +0000 Subject: [PATCH 1016/1291] chore(analysis/special_functions/pow): move around rpow/cpow tactics (#19022) This PR splits up the code that implements `norm_num` and `positivity` for real / complex powers, and puts it into the files where those operations are defined, for easier discoverability. --- src/analysis/fourier/poisson_summation.lean | 1 - .../special_functions/japanese_bracket.lean | 1 - .../special_functions/pow/complex.lean | 52 +++++ .../special_functions/pow/nnreal.lean | 89 ++++++++ src/analysis/special_functions/pow/real.lean | 76 +++++++ .../special_functions/pow/tactic.lean | 213 ------------------ src/combinatorics/additive/behrend.lean | 1 - .../simple_graph/regularity/bound.lean | 2 +- src/number_theory/bertrand.lean | 1 - test/norm_num_ext.lean | 2 +- test/positivity.lean | 2 +- 11 files changed, 220 insertions(+), 220 deletions(-) delete mode 100644 src/analysis/special_functions/pow/tactic.lean diff --git a/src/analysis/fourier/poisson_summation.lean b/src/analysis/fourier/poisson_summation.lean index 682068c630569..17e5fae7dc9cc 100644 --- a/src/analysis/fourier/poisson_summation.lean +++ b/src/analysis/fourier/poisson_summation.lean @@ -7,7 +7,6 @@ import analysis.fourier.add_circle import analysis.fourier.fourier_transform import analysis.p_series import analysis.schwartz_space -import analysis.special_functions.pow.tactic /-! # Poisson's summation formula diff --git a/src/analysis/special_functions/japanese_bracket.lean b/src/analysis/special_functions/japanese_bracket.lean index 6b6808010278e..d287916d88f20 100644 --- a/src/analysis/special_functions/japanese_bracket.lean +++ b/src/analysis/special_functions/japanese_bracket.lean @@ -5,7 +5,6 @@ Authors: Moritz Doll -/ import measure_theory.measure.haar_lebesgue import measure_theory.integral.layercake -import analysis.special_functions.pow.tactic /-! # Japanese Bracket diff --git a/src/analysis/special_functions/pow/complex.lean b/src/analysis/special_functions/pow/complex.lean index d60d119b4a50b..d8a42ea6f78a9 100644 --- a/src/analysis/special_functions/pow/complex.lean +++ b/src/analysis/special_functions/pow/complex.lean @@ -177,3 +177,55 @@ lemma cpow_conj (x : ℂ) (n : ℂ) (hx : x.arg ≠ π) : x ^ conj n = conj (con by rw [conj_cpow _ _ hx, conj_conj] end complex + +section tactics +/-! +## Tactic extensions for complex powers +-/ + +namespace norm_num + +theorem cpow_pos (a b : ℂ) (b' : ℕ) (c : ℂ) (hb : b = b') (h : a ^ b' = c) : a ^ b = c := +by rw [← h, hb, complex.cpow_nat_cast] + +theorem cpow_neg (a b : ℂ) (b' : ℕ) (c c' : ℂ) + (hb : b = b') (h : a ^ b' = c) (hc : c⁻¹ = c') : a ^ -b = c' := +by rw [← hc, ← h, hb, complex.cpow_neg, complex.cpow_nat_cast] + +open tactic + +/-- Generalized version of `prove_cpow`, `prove_nnrpow`, `prove_ennrpow`. -/ +meta def prove_rpow' (pos neg zero : name) (α β one a b : expr) : tactic (expr × expr) := do + na ← a.to_rat, + icα ← mk_instance_cache α, + icβ ← mk_instance_cache β, + match match_sign b with + | sum.inl b := do + nc ← mk_instance_cache `(ℕ), + (icβ, nc, b', hb) ← prove_nat_uncast icβ nc b, + (icα, c, h) ← prove_pow a na icα b', + cr ← c.to_rat, + (icα, c', hc) ← prove_inv icα c cr, + pure (c', (expr.const neg []).mk_app [a, b, b', c, c', hb, h, hc]) + | sum.inr ff := pure (one, expr.const zero [] a) + | sum.inr tt := do + nc ← mk_instance_cache `(ℕ), + (icβ, nc, b', hb) ← prove_nat_uncast icβ nc b, + (icα, c, h) ← prove_pow a na icα b', + pure (c, (expr.const pos []).mk_app [a, b, b', c, hb, h]) + end + +/-- Evaluate `complex.cpow a b` where `a` is a rational numeral and `b` is an integer. -/ +meta def prove_cpow : expr → expr → tactic (expr × expr) := +prove_rpow' ``cpow_pos ``cpow_neg ``complex.cpow_zero `(ℂ) `(ℂ) `(1:ℂ) + +/-- Evaluates expressions of the form `cpow a b` and `a ^ b` in the special case where +`b` is an integer and `a` is a positive rational (so it's really just a rational power). -/ +@[norm_num] meta def eval_cpow : expr → tactic (expr × expr) +| `(@has_pow.pow _ _ complex.has_pow %%a %%b) := b.to_int >> prove_cpow a b +| `(complex.cpow %%a %%b) := b.to_int >> prove_cpow a b +| _ := tactic.failed + +end norm_num + +end tactics diff --git a/src/analysis/special_functions/pow/nnreal.lean b/src/analysis/special_functions/pow/nnreal.lean index 155a69cd6c534..65aec15568256 100644 --- a/src/analysis/special_functions/pow/nnreal.lean +++ b/src/analysis/special_functions/pow/nnreal.lean @@ -691,3 +691,92 @@ lemma rpow_left_bijective {x : ℝ} (hx : x ≠ 0) : ⟨rpow_left_injective hx, rpow_left_surjective hx⟩ end ennreal + +section tactics +/-! +## Tactic extensions for powers on `ℝ≥0` and `ℝ≥0∞` +-/ + +namespace norm_num + +theorem nnrpow_pos (a : ℝ≥0) (b : ℝ) (b' : ℕ) (c : ℝ≥0) + (hb : b = b') (h : a ^ b' = c) : a ^ b = c := +by rw [← h, hb, nnreal.rpow_nat_cast] + +theorem nnrpow_neg (a : ℝ≥0) (b : ℝ) (b' : ℕ) (c c' : ℝ≥0) + (hb : b = b') (h : a ^ b' = c) (hc : c⁻¹ = c') : a ^ -b = c' := +by rw [← hc, ← h, hb, nnreal.rpow_neg, nnreal.rpow_nat_cast] + +theorem ennrpow_pos (a : ℝ≥0∞) (b : ℝ) (b' : ℕ) (c : ℝ≥0∞) + (hb : b = b') (h : a ^ b' = c) : a ^ b = c := +by rw [← h, hb, ennreal.rpow_nat_cast] + +theorem ennrpow_neg (a : ℝ≥0∞) (b : ℝ) (b' : ℕ) (c c' : ℝ≥0∞) + (hb : b = b') (h : a ^ b' = c) (hc : c⁻¹ = c') : a ^ -b = c' := +by rw [← hc, ← h, hb, ennreal.rpow_neg, ennreal.rpow_nat_cast] + +/-- Evaluate `nnreal.rpow a b` where `a` is a rational numeral and `b` is an integer. -/ +meta def prove_nnrpow : expr → expr → tactic (expr × expr) := +prove_rpow' ``nnrpow_pos ``nnrpow_neg ``nnreal.rpow_zero `(ℝ≥0) `(ℝ) `(1:ℝ≥0) + +/-- Evaluate `ennreal.rpow a b` where `a` is a rational numeral and `b` is an integer. -/ +meta def prove_ennrpow : expr → expr → tactic (expr × expr) := +prove_rpow' ``ennrpow_pos ``ennrpow_neg ``ennreal.rpow_zero `(ℝ≥0∞) `(ℝ) `(1:ℝ≥0∞) + +/-- Evaluates expressions of the form `rpow a b` and `a ^ b` in the special case where +`b` is an integer and `a` is a positive rational (so it's really just a rational power). -/ +@[norm_num] meta def eval_nnrpow_ennrpow : expr → tactic (expr × expr) +| `(@has_pow.pow _ _ nnreal.real.has_pow %%a %%b) := b.to_int >> prove_nnrpow a b +| `(nnreal.rpow %%a %%b) := b.to_int >> prove_nnrpow a b +| `(@has_pow.pow _ _ ennreal.real.has_pow %%a %%b) := b.to_int >> prove_ennrpow a b +| `(ennreal.rpow %%a %%b) := b.to_int >> prove_ennrpow a b +| _ := tactic.failed + +end norm_num + +namespace tactic +namespace positivity + +private lemma nnrpow_pos {a : ℝ≥0} (ha : 0 < a) (b : ℝ) : 0 < a ^ b := nnreal.rpow_pos ha + +/-- Auxiliary definition for the `positivity` tactic to handle real powers of nonnegative reals. -/ +meta def prove_nnrpow (a b : expr) : tactic strictness := +do + strictness_a ← core a, + match strictness_a with + | positive p := positive <$> mk_app ``nnrpow_pos [p, b] + | _ := failed -- We already know `0 ≤ x` for all `x : ℝ≥0` + end + +private lemma ennrpow_pos {a : ℝ≥0∞} {b : ℝ} (ha : 0 < a) (hb : 0 < b) : 0 < a ^ b := +ennreal.rpow_pos_of_nonneg ha hb.le + +/-- Auxiliary definition for the `positivity` tactic to handle real powers of extended nonnegative +reals. -/ +meta def prove_ennrpow (a b : expr) : tactic strictness := +do + strictness_a ← core a, + strictness_b ← core b, + match strictness_a, strictness_b with + | positive pa, positive pb := positive <$> mk_app ``ennrpow_pos [pa, pb] + | positive pa, nonnegative pb := positive <$> mk_app ``ennreal.rpow_pos_of_nonneg [pa, pb] + | _, _ := failed -- We already know `0 ≤ x` for all `x : ℝ≥0∞` + end + +end positivity + +open positivity + +/-- Extension for the `positivity` tactic: exponentiation by a real number is nonnegative when the +base is nonnegative and positive when the base is positive. -/ +@[positivity] +meta def positivity_nnrpow_ennrpow : expr → tactic strictness +| `(@has_pow.pow _ _ nnreal.real.has_pow %%a %%b) := prove_nnrpow a b +| `(nnreal.rpow %%a %%b) := prove_nnrpow a b +| `(@has_pow.pow _ _ ennreal.real.has_pow %%a %%b) := prove_ennrpow a b +| `(ennreal.rpow %%a %%b) := prove_ennrpow a b +| _ := failed + +end tactic + +end tactics diff --git a/src/analysis/special_functions/pow/real.lean b/src/analysis/special_functions/pow/real.lean index c938399127f70..6815e879fee35 100644 --- a/src/analysis/special_functions/pow/real.lean +++ b/src/analysis/special_functions/pow/real.lean @@ -639,3 +639,79 @@ begin end end real + +section tactics +/-! +## Tactic extensions for real powers +-/ + +namespace norm_num +open tactic + +theorem rpow_pos (a b : ℝ) (b' : ℕ) (c : ℝ) (hb : (b':ℝ) = b) (h : a ^ b' = c) : a ^ b = c := +by rw [← h, ← hb, real.rpow_nat_cast] + +theorem rpow_neg (a b : ℝ) (b' : ℕ) (c c' : ℝ) + (a0 : 0 ≤ a) (hb : (b':ℝ) = b) (h : a ^ b' = c) (hc : c⁻¹ = c') : a ^ -b = c' := +by rw [← hc, ← h, ← hb, real.rpow_neg a0, real.rpow_nat_cast] + +/-- Evaluate `real.rpow a b` where `a` is a rational numeral and `b` is an integer. +(This cannot go via the generalized version `prove_rpow'` because `rpow_pos` has a side condition; +we do not attempt to evaluate `a ^ b` where `a` and `b` are both negative because it comes +out to some garbage.) -/ +meta def prove_rpow (a b : expr) : tactic (expr × expr) := do + na ← a.to_rat, + ic ← mk_instance_cache `(ℝ), + match match_sign b with + | sum.inl b := do + (ic, a0) ← guard (na ≥ 0) >> prove_nonneg ic a, + nc ← mk_instance_cache `(ℕ), + (ic, nc, b', hb) ← prove_nat_uncast ic nc b, + (ic, c, h) ← prove_pow a na ic b', + cr ← c.to_rat, + (ic, c', hc) ← prove_inv ic c cr, + pure (c', (expr.const ``rpow_neg []).mk_app [a, b, b', c, c', a0, hb, h, hc]) + | sum.inr ff := pure (`(1:ℝ), expr.const ``real.rpow_zero [] a) + | sum.inr tt := do + nc ← mk_instance_cache `(ℕ), + (ic, nc, b', hb) ← prove_nat_uncast ic nc b, + (ic, c, h) ← prove_pow a na ic b', + pure (c, (expr.const ``rpow_pos []).mk_app [a, b, b', c, hb, h]) + end + +/-- Evaluates expressions of the form `rpow a b` and `a ^ b` in the special case where +`b` is an integer and `a` is a positive rational (so it's really just a rational power). -/ +@[norm_num] meta def eval_rpow : expr → tactic (expr × expr) +| `(@has_pow.pow _ _ real.has_pow %%a %%b) := b.to_int >> prove_rpow a b +| `(real.rpow %%a %%b) := b.to_int >> prove_rpow a b +| _ := tactic.failed +end norm_num + +namespace tactic +namespace positivity + +/-- Auxiliary definition for the `positivity` tactic to handle real powers of reals. -/ +meta def prove_rpow (a b : expr) : tactic strictness := +do + strictness_a ← core a, + match strictness_a with + | nonnegative p := nonnegative <$> mk_app ``real.rpow_nonneg_of_nonneg [p, b] + | positive p := positive <$> mk_app ``real.rpow_pos_of_pos [p, b] + | _ := failed + end + +end positivity + +open positivity + +/-- Extension for the `positivity` tactic: exponentiation by a real number is nonnegative when the +base is nonnegative and positive when the base is positive. -/ +@[positivity] +meta def positivity_rpow : expr → tactic strictness +| `(@has_pow.pow _ _ real.has_pow %%a %%b) := prove_rpow a b +| `(real.rpow %%a %%b) := prove_rpow a b +| _ := failed + +end tactic + +end tactics diff --git a/src/analysis/special_functions/pow/tactic.lean b/src/analysis/special_functions/pow/tactic.lean deleted file mode 100644 index 1a426e55b07c8..0000000000000 --- a/src/analysis/special_functions/pow/tactic.lean +++ /dev/null @@ -1,213 +0,0 @@ -/- -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, Sébastien Gouëzel, - Rémy Degenne, David Loeffler --/ -import analysis.special_functions.pow.nnreal - -/-! -# Tactics for power functions - -This file contains extensions to the `norm_num` and `positivity` tactics to handle power operations -on `ℂ`, `ℝ`, `ℝ≥0`, and `ℝ≥0∞`. - -TODO: Split up the contents of this file and merge with other files in -`analysis/special_functions/pow/`, to keep the tactics together with the corresponding definitions. - --/ - -open_locale nnreal ennreal - -/-! -## Complex case --/ - -namespace norm_num - -theorem cpow_pos (a b : ℂ) (b' : ℕ) (c : ℂ) (hb : b = b') (h : a ^ b' = c) : a ^ b = c := -by rw [← h, hb, complex.cpow_nat_cast] - -theorem cpow_neg (a b : ℂ) (b' : ℕ) (c c' : ℂ) - (hb : b = b') (h : a ^ b' = c) (hc : c⁻¹ = c') : a ^ -b = c' := -by rw [← hc, ← h, hb, complex.cpow_neg, complex.cpow_nat_cast] - -end norm_num - -/-! -## Real case --/ - -namespace norm_num -open tactic - -theorem rpow_pos (a b : ℝ) (b' : ℕ) (c : ℝ) (hb : (b':ℝ) = b) (h : a ^ b' = c) : a ^ b = c := -by rw [← h, ← hb, real.rpow_nat_cast] - -theorem rpow_neg (a b : ℝ) (b' : ℕ) (c c' : ℝ) - (a0 : 0 ≤ a) (hb : (b':ℝ) = b) (h : a ^ b' = c) (hc : c⁻¹ = c') : a ^ -b = c' := -by rw [← hc, ← h, ← hb, real.rpow_neg a0, real.rpow_nat_cast] - -/-- Evaluate `real.rpow a b` where `a` is a rational numeral and `b` is an integer. -(This cannot go via the generalized version `prove_rpow'` because `rpow_pos` has a side condition; -we do not attempt to evaluate `a ^ b` where `a` and `b` are both negative because it comes -out to some garbage.) -/ -meta def prove_rpow (a b : expr) : tactic (expr × expr) := do - na ← a.to_rat, - ic ← mk_instance_cache `(ℝ), - match match_sign b with - | sum.inl b := do - (ic, a0) ← guard (na ≥ 0) >> prove_nonneg ic a, - nc ← mk_instance_cache `(ℕ), - (ic, nc, b', hb) ← prove_nat_uncast ic nc b, - (ic, c, h) ← prove_pow a na ic b', - cr ← c.to_rat, - (ic, c', hc) ← prove_inv ic c cr, - pure (c', (expr.const ``rpow_neg []).mk_app [a, b, b', c, c', a0, hb, h, hc]) - | sum.inr ff := pure (`(1:ℝ), expr.const ``real.rpow_zero [] a) - | sum.inr tt := do - nc ← mk_instance_cache `(ℕ), - (ic, nc, b', hb) ← prove_nat_uncast ic nc b, - (ic, c, h) ← prove_pow a na ic b', - pure (c, (expr.const ``rpow_pos []).mk_app [a, b, b', c, hb, h]) - end - -end norm_num - -namespace tactic -namespace positivity - -/-- Auxiliary definition for the `positivity` tactic to handle real powers of reals. -/ -meta def prove_rpow (a b : expr) : tactic strictness := -do - strictness_a ← core a, - match strictness_a with - | nonnegative p := nonnegative <$> mk_app ``real.rpow_nonneg_of_nonneg [p, b] - | positive p := positive <$> mk_app ``real.rpow_pos_of_pos [p, b] - | _ := failed - end - -end positivity -end tactic - -/-! -## General-purpose tactics - -The following code covers all the cases `ℂ`, `ℝ`, `ℝ≥0`, and `ℝ≥0∞` together. --/ - -namespace norm_num -open tactic - -/-- Generalized version of `prove_cpow`, `prove_nnrpow`, `prove_ennrpow`. -/ -meta def prove_rpow' (pos neg zero : name) (α β one a b : expr) : tactic (expr × expr) := do - na ← a.to_rat, - icα ← mk_instance_cache α, - icβ ← mk_instance_cache β, - match match_sign b with - | sum.inl b := do - nc ← mk_instance_cache `(ℕ), - (icβ, nc, b', hb) ← prove_nat_uncast icβ nc b, - (icα, c, h) ← prove_pow a na icα b', - cr ← c.to_rat, - (icα, c', hc) ← prove_inv icα c cr, - pure (c', (expr.const neg []).mk_app [a, b, b', c, c', hb, h, hc]) - | sum.inr ff := pure (one, expr.const zero [] a) - | sum.inr tt := do - nc ← mk_instance_cache `(ℕ), - (icβ, nc, b', hb) ← prove_nat_uncast icβ nc b, - (icα, c, h) ← prove_pow a na icα b', - pure (c, (expr.const pos []).mk_app [a, b, b', c, hb, h]) - end - -open_locale nnreal ennreal - -theorem nnrpow_pos (a : ℝ≥0) (b : ℝ) (b' : ℕ) (c : ℝ≥0) - (hb : b = b') (h : a ^ b' = c) : a ^ b = c := -by rw [← h, hb, nnreal.rpow_nat_cast] - -theorem nnrpow_neg (a : ℝ≥0) (b : ℝ) (b' : ℕ) (c c' : ℝ≥0) - (hb : b = b') (h : a ^ b' = c) (hc : c⁻¹ = c') : a ^ -b = c' := -by rw [← hc, ← h, hb, nnreal.rpow_neg, nnreal.rpow_nat_cast] - -theorem ennrpow_pos (a : ℝ≥0∞) (b : ℝ) (b' : ℕ) (c : ℝ≥0∞) - (hb : b = b') (h : a ^ b' = c) : a ^ b = c := -by rw [← h, hb, ennreal.rpow_nat_cast] - -theorem ennrpow_neg (a : ℝ≥0∞) (b : ℝ) (b' : ℕ) (c c' : ℝ≥0∞) - (hb : b = b') (h : a ^ b' = c) (hc : c⁻¹ = c') : a ^ -b = c' := -by rw [← hc, ← h, hb, ennreal.rpow_neg, ennreal.rpow_nat_cast] - -/-- Evaluate `complex.cpow a b` where `a` is a rational numeral and `b` is an integer. -/ -meta def prove_cpow : expr → expr → tactic (expr × expr) := -prove_rpow' ``cpow_pos ``cpow_neg ``complex.cpow_zero `(ℂ) `(ℂ) `(1:ℂ) - -/-- Evaluate `nnreal.rpow a b` where `a` is a rational numeral and `b` is an integer. -/ -meta def prove_nnrpow : expr → expr → tactic (expr × expr) := -prove_rpow' ``nnrpow_pos ``nnrpow_neg ``nnreal.rpow_zero `(ℝ≥0) `(ℝ) `(1:ℝ≥0) - -/-- Evaluate `ennreal.rpow a b` where `a` is a rational numeral and `b` is an integer. -/ -meta def prove_ennrpow : expr → expr → tactic (expr × expr) := -prove_rpow' ``ennrpow_pos ``ennrpow_neg ``ennreal.rpow_zero `(ℝ≥0∞) `(ℝ) `(1:ℝ≥0∞) - -/-- Evaluates expressions of the form `rpow a b`, `cpow a b` and `a ^ b` in the special case where -`b` is an integer and `a` is a positive rational (so it's really just a rational power). -/ -@[norm_num] meta def eval_rpow_cpow : expr → tactic (expr × expr) -| `(@has_pow.pow _ _ real.has_pow %%a %%b) := b.to_int >> prove_rpow a b -| `(real.rpow %%a %%b) := b.to_int >> prove_rpow a b -| `(@has_pow.pow _ _ complex.has_pow %%a %%b) := b.to_int >> prove_cpow a b -| `(complex.cpow %%a %%b) := b.to_int >> prove_cpow a b -| `(@has_pow.pow _ _ nnreal.real.has_pow %%a %%b) := b.to_int >> prove_nnrpow a b -| `(nnreal.rpow %%a %%b) := b.to_int >> prove_nnrpow a b -| `(@has_pow.pow _ _ ennreal.real.has_pow %%a %%b) := b.to_int >> prove_ennrpow a b -| `(ennreal.rpow %%a %%b) := b.to_int >> prove_ennrpow a b -| _ := tactic.failed - -end norm_num - -namespace tactic -namespace positivity - -private lemma nnrpow_pos {a : ℝ≥0} (ha : 0 < a) (b : ℝ) : 0 < a ^ b := nnreal.rpow_pos ha - -/-- Auxiliary definition for the `positivity` tactic to handle real powers of nonnegative reals. -/ -meta def prove_nnrpow (a b : expr) : tactic strictness := -do - strictness_a ← core a, - match strictness_a with - | positive p := positive <$> mk_app ``nnrpow_pos [p, b] - | _ := failed -- We already know `0 ≤ x` for all `x : ℝ≥0` - end - -private lemma ennrpow_pos {a : ℝ≥0∞} {b : ℝ} (ha : 0 < a) (hb : 0 < b) : 0 < a ^ b := -ennreal.rpow_pos_of_nonneg ha hb.le - -/-- Auxiliary definition for the `positivity` tactic to handle real powers of extended nonnegative -reals. -/ -meta def prove_ennrpow (a b : expr) : tactic strictness := -do - strictness_a ← core a, - strictness_b ← core b, - match strictness_a, strictness_b with - | positive pa, positive pb := positive <$> mk_app ``ennrpow_pos [pa, pb] - | positive pa, nonnegative pb := positive <$> mk_app ``ennreal.rpow_pos_of_nonneg [pa, pb] - | _, _ := failed -- We already know `0 ≤ x` for all `x : ℝ≥0∞` - end - -end positivity - -open positivity - -/-- Extension for the `positivity` tactic: exponentiation by a real number is nonnegative when the -base is nonnegative and positive when the base is positive. -/ -@[positivity] -meta def positivity_rpow : expr → tactic strictness -| `(@has_pow.pow _ _ real.has_pow %%a %%b) := prove_rpow a b -| `(real.rpow %%a %%b) := prove_rpow a b -| `(@has_pow.pow _ _ nnreal.real.has_pow %%a %%b) := prove_nnrpow a b -| `(nnreal.rpow %%a %%b) := prove_nnrpow a b -| `(@has_pow.pow _ _ ennreal.real.has_pow %%a %%b) := prove_ennrpow a b -| `(ennreal.rpow %%a %%b) := prove_ennrpow a b -| _ := failed - -end tactic diff --git a/src/combinatorics/additive/behrend.lean b/src/combinatorics/additive/behrend.lean index 6e4eae183ca42..857a123b6c14e 100644 --- a/src/combinatorics/additive/behrend.lean +++ b/src/combinatorics/additive/behrend.lean @@ -7,7 +7,6 @@ import analysis.inner_product_space.pi_L2 import combinatorics.additive.salem_spencer import combinatorics.pigeonhole import data.complex.exponential_bounds -import analysis.special_functions.pow.tactic /-! # Behrend's bound on Roth numbers diff --git a/src/combinatorics/simple_graph/regularity/bound.lean b/src/combinatorics/simple_graph/regularity/bound.lean index 756d5fb7b80ea..41e723527b1c1 100644 --- a/src/combinatorics/simple_graph/regularity/bound.lean +++ b/src/combinatorics/simple_graph/regularity/bound.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies, Bhavik Mehta -/ import algebra.order.chebyshev -import analysis.special_functions.pow.tactic +import analysis.special_functions.pow.real import order.partition.equipartition /-! diff --git a/src/number_theory/bertrand.lean b/src/number_theory/bertrand.lean index b98486b53e1d0..b9dbb2af60d55 100644 --- a/src/number_theory/bertrand.lean +++ b/src/number_theory/bertrand.lean @@ -7,7 +7,6 @@ import data.nat.choose.factorization import data.nat.prime_norm_num import number_theory.primorial import analysis.convex.specific_functions -import analysis.special_functions.pow.tactic /-! # Bertrand's Postulate diff --git a/test/norm_num_ext.lean b/test/norm_num_ext.lean index e1f93a9210b86..96648a919b00d 100644 --- a/test/norm_num_ext.lean +++ b/test/norm_num_ext.lean @@ -9,7 +9,7 @@ import data.int.gcd import data.nat.fib import data.nat.prime import data.nat.sqrt_norm_num -import analysis.special_functions.pow.tactic +import analysis.special_functions.pow.real import number_theory.legendre_symbol.norm_num /-! diff --git a/test/positivity.lean b/test/positivity.lean index 30b644d798dab..f4dc7425036ef 100644 --- a/test/positivity.lean +++ b/test/positivity.lean @@ -1,7 +1,7 @@ import algebra.order.interval import algebra.order.smul import analysis.normed.group.basic -import analysis.special_functions.pow.tactic +import analysis.special_functions.pow.nnreal import combinatorics.simple_graph.density import data.complex.exponential import data.rat.nnrat From b76e9f654df09f8a832aeee712511fe5f3e57869 Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Tue, 16 May 2023 05:07:51 +0000 Subject: [PATCH 1017/1291] feat(analysis/special_functions/gamma): holomorphy of `1 / Gamma` (#19012) This PR makes two changes to the Gamma function code: - drastically shorten the proof of differentiability of the Gamma function by applying general results on Mellin transforms; - add the fact that `1 / Gamma` is differentiable everywhere (including at the poles of the Gamma function). --- .../special_functions/gamma/basic.lean | 159 +++--------------- .../special_functions/gamma/beta.lean | 45 +++++ 2 files changed, 67 insertions(+), 137 deletions(-) diff --git a/src/analysis/special_functions/gamma/basic.lean b/src/analysis/special_functions/gamma/basic.lean index 1d488c767c51c..e496adcd4030d 100644 --- a/src/analysis/special_functions/gamma/basic.lean +++ b/src/analysis/special_functions/gamma/basic.lean @@ -5,7 +5,7 @@ Authors: David Loeffler -/ import measure_theory.integral.exp_decay import analysis.special_functions.improper_integrals -import analysis.calculus.parametric_integral +import analysis.mellin_transform /-! # The Gamma function @@ -382,152 +382,37 @@ end end Gamma_def -end complex - /-! Now check that the `Γ` function is differentiable, wherever this makes sense. -/ section Gamma_has_deriv -/-- Integrand for the derivative of the `Γ` function -/ -def dGamma_integrand (s : ℂ) (x : ℝ) : ℂ := exp (-x) * log x * x ^ (s - 1) - -/-- Integrand for the absolute value of the derivative of the `Γ` function -/ -def dGamma_integrand_real (s x : ℝ) : ℝ := |exp (-x) * log x * x ^ (s - 1)| - -lemma dGamma_integrand_is_o_at_top (s : ℝ) : - (λ x : ℝ, exp (-x) * log x * x ^ (s - 1)) =o[at_top] (λ x, exp (-(1/2) * x)) := -begin - refine is_o_of_tendsto (λ x hx, _) _, - { exfalso, exact (-(1/2) * x).exp_pos.ne' hx, }, - have : eventually_eq at_top (λ (x : ℝ), exp (-x) * log x * x ^ (s - 1) / exp (-(1 / 2) * x)) - (λ (x : ℝ), (λ z:ℝ, exp (1 / 2 * z) / z ^ s) x * (λ z:ℝ, z / log z) x)⁻¹, - { refine eventually_of_mem (Ioi_mem_at_top 1) _, - intros x hx, dsimp, - replace hx := lt_trans zero_lt_one (mem_Ioi.mp hx), - rw [real.exp_neg, neg_mul, real.exp_neg, rpow_sub hx], - have : exp x = exp(x/2) * exp(x/2), - { rw [←real.exp_add, add_halves], }, - rw this, field_simp [hx.ne', exp_ne_zero (x/2)], ring, }, - refine tendsto.congr' this.symm (tendsto.inv_tendsto_at_top _), - apply tendsto.at_top_mul_at_top (tendsto_exp_mul_div_rpow_at_top s (1/2) one_half_pos), - refine tendsto.congr' _ ((tendsto_exp_div_pow_at_top 1).comp tendsto_log_at_top), - apply eventually_eq_of_mem (Ioi_mem_at_top (0:ℝ)), - intros x hx, simp [exp_log hx], -end - -/-- Absolute convergence of the integral which will give the derivative of the `Γ` function on -`1 < re s`. -/ -lemma dGamma_integral_abs_convergent (s : ℝ) (hs : 1 < s) : - integrable_on (λ x:ℝ, ‖exp (-x) * log x * x ^ (s-1)‖) (Ioi 0) := -begin - rw [←Ioc_union_Ioi_eq_Ioi (@zero_le_one ℝ _ _ _ _), integrable_on_union], - refine ⟨⟨_, _⟩, _⟩, - { refine continuous_on.ae_strongly_measurable (continuous_on.mul _ _).norm measurable_set_Ioc, - { refine (continuous_exp.comp continuous_neg).continuous_on.mul (continuous_on_log.mono _), - simp, }, - { apply continuous_on_id.rpow_const, intros x hx, right, linarith }, }, - { apply has_finite_integral_of_bounded, - swap, { exact 1 / (s - 1), }, - refine (ae_restrict_iff' measurable_set_Ioc).mpr (ae_of_all _ (λ x hx, _)), - rw [norm_norm, norm_eq_abs, mul_assoc, abs_mul, ←one_mul (1 / (s - 1))], - refine mul_le_mul _ _ (abs_nonneg _) zero_le_one, - { rw [abs_of_pos (exp_pos(-x)), exp_le_one_iff, neg_le, neg_zero], exact hx.1.le }, - { exact (abs_log_mul_self_rpow_lt x (s-1) hx.1 hx.2 (sub_pos.mpr hs)).le }, }, - { have := (dGamma_integrand_is_o_at_top s).is_O.norm_left, - refine integrable_of_is_O_exp_neg one_half_pos (continuous_on.mul _ _).norm this, - { refine (continuous_exp.comp continuous_neg).continuous_on.mul (continuous_on_log.mono _), - simp, }, - { apply continuous_at.continuous_on (λ x hx, _), - apply continuous_at_id.rpow continuous_at_const, - dsimp, right, linarith, }, } -end - -/-- A uniform bound for the `s`-derivative of the `Γ` integrand for `s` in vertical strips. -/ -lemma loc_unif_bound_dGamma_integrand {t : ℂ} {s1 s2 x : ℝ} (ht1 : s1 ≤ t.re) - (ht2: t.re ≤ s2) (hx : 0 < x) : - ‖dGamma_integrand t x‖ ≤ dGamma_integrand_real s1 x + dGamma_integrand_real s2 x := -begin - rcases le_or_lt 1 x with h|h, - { -- case 1 ≤ x - refine le_add_of_nonneg_of_le (abs_nonneg _) _, - rw [dGamma_integrand, dGamma_integrand_real, complex.norm_eq_abs, map_mul, abs_mul, - ←complex.of_real_mul, complex.abs_of_real], - refine mul_le_mul_of_nonneg_left _ (abs_nonneg _), - rw complex.abs_cpow_eq_rpow_re_of_pos hx, - refine le_trans _ (le_abs_self _), - apply rpow_le_rpow_of_exponent_le h, - rw [complex.sub_re, complex.one_re], linarith, }, - { refine le_add_of_le_of_nonneg _ (abs_nonneg _), - rw [dGamma_integrand, dGamma_integrand_real, complex.norm_eq_abs, map_mul, abs_mul, - ←complex.of_real_mul, complex.abs_of_real], - refine mul_le_mul_of_nonneg_left _ (abs_nonneg _), - rw complex.abs_cpow_eq_rpow_re_of_pos hx, - refine le_trans _ (le_abs_self _), - apply rpow_le_rpow_of_exponent_ge hx h.le, - rw [complex.sub_re, complex.one_re], linarith, }, -end - -namespace complex +/-- Rewrite the Gamma integral as an example of a Mellin transform. -/ +lemma Gamma_integral_eq_mellin : Gamma_integral = mellin (λ x, real.exp (-x)) := +funext (λ s, by simp only [mellin, Gamma_integral, smul_eq_mul, mul_comm]) -/-- The derivative of the `Γ` integral, at any `s ∈ ℂ` with `1 < re s`, is given by the integral -of `exp (-x) * log x * x ^ (s - 1)` over `[0, ∞)`. -/ -theorem has_deriv_at_Gamma_integral {s : ℂ} (hs : 1 < s.re) : - (integrable_on (λ x, real.exp (-x) * real.log x * x ^ (s - 1) : ℝ → ℂ) (Ioi 0) volume) ∧ - (has_deriv_at Gamma_integral (∫ x:ℝ in Ioi 0, real.exp (-x) * real.log x * x ^ (s - 1)) s) := +/-- The derivative of the `Γ` integral, at any `s ∈ ℂ` with `1 < re s`, is given by the Melllin +transform of `log t * exp (-t)`. -/ +theorem has_deriv_at_Gamma_integral {s : ℂ} (hs : 0 < s.re) : + has_deriv_at Gamma_integral (∫ (t : ℝ) in Ioi 0, t ^ (s - 1) * (real.log t * real.exp (-t))) s := begin - let ε := (s.re - 1) / 2, - let μ := volume.restrict (Ioi (0:ℝ)), - let bound := (λ x:ℝ, dGamma_integrand_real (s.re - ε) x + dGamma_integrand_real (s.re + ε) x), - have cont : ∀ (t : ℂ), continuous_on (λ x, real.exp (-x) * x ^ (t - 1) : ℝ → ℂ) (Ioi 0), - { intro t, apply (continuous_of_real.comp continuous_neg.exp).continuous_on.mul, - apply continuous_at.continuous_on, intros x hx, - refine (continuous_at_cpow_const _).comp continuous_of_real.continuous_at, - exact or.inl hx, }, - have eps_pos: 0 < ε := div_pos (sub_pos.mpr hs) zero_lt_two, - have hF_meas : ∀ᶠ (t : ℂ) in 𝓝 s, - ae_strongly_measurable (λ x, real.exp(-x) * x ^ (t - 1) : ℝ → ℂ) μ, - { apply eventually_of_forall, intro t, - exact (cont t).ae_strongly_measurable measurable_set_Ioi, }, - have hF'_meas : ae_strongly_measurable (dGamma_integrand s) μ, - { refine continuous_on.ae_strongly_measurable _ measurable_set_Ioi, - have : dGamma_integrand s = (λ x, real.exp (-x) * x ^ (s - 1) * real.log x : ℝ → ℂ), - { ext1, simp only [dGamma_integrand], ring }, - rw this, - refine continuous_on.mul (cont s) (continuous_at.continuous_on _), - exact λ x hx, continuous_of_real.continuous_at.comp (continuous_at_log (mem_Ioi.mp hx).ne'), }, - have h_bound : ∀ᵐ (x : ℝ) ∂μ, ∀ (t : ℂ), t ∈ metric.ball s ε → ‖ dGamma_integrand t x ‖ ≤ bound x, - { refine (ae_restrict_iff' measurable_set_Ioi).mpr (ae_of_all _ (λ x hx, _)), - intros t ht, - rw [metric.mem_ball, complex.dist_eq] at ht, - replace ht := lt_of_le_of_lt (complex.abs_re_le_abs $ t - s ) ht, - rw [complex.sub_re, @abs_sub_lt_iff ℝ _ t.re s.re ((s.re - 1) / 2) ] at ht, - refine loc_unif_bound_dGamma_integrand _ _ hx, - all_goals { simp only [ε], linarith } }, - have bound_integrable : integrable bound μ, - { apply integrable.add, - { refine dGamma_integral_abs_convergent (s.re - ε) _, - field_simp, rw one_lt_div, - { linarith }, { exact zero_lt_two }, }, - { refine dGamma_integral_abs_convergent (s.re + ε) _, linarith, }, }, - have h_diff : ∀ᵐ (x : ℝ) ∂μ, ∀ (t : ℂ), t ∈ metric.ball s ε - → has_deriv_at (λ u, real.exp (-x) * x ^ (u - 1) : ℂ → ℂ) (dGamma_integrand t x) t, - { refine (ae_restrict_iff' measurable_set_Ioi).mpr (ae_of_all _ (λ x hx, _)), - intros t ht, rw mem_Ioi at hx, - simp only [dGamma_integrand], - rw mul_assoc, - apply has_deriv_at.const_mul, - rw [of_real_log hx.le, mul_comm], - have := ((has_deriv_at_id t).sub_const 1).const_cpow (or.inl (of_real_ne_zero.mpr hx.ne')), - rwa mul_one at this }, - exact (has_deriv_at_integral_of_dominated_loc_of_deriv_le eps_pos hF_meas - (Gamma_integral_convergent (zero_lt_one.trans hs)) hF'_meas h_bound bound_integrable h_diff), + rw Gamma_integral_eq_mellin, + convert mellin_has_deriv_of_is_O_rpow _ _ (lt_add_one _) _ hs, + { refine (continuous.continuous_on _).locally_integrable_on measurable_set_Ioi, + exact continuous_of_real.comp (real.continuous_exp.comp continuous_neg), }, + { rw [←is_O_norm_left], + simp_rw [complex.norm_eq_abs, abs_of_real, ←real.norm_eq_abs, is_O_norm_left], + simpa only [neg_one_mul] using (is_o_exp_neg_mul_rpow_at_top zero_lt_one _).is_O }, + { simp_rw [neg_zero, rpow_zero], + refine is_O_const_of_tendsto (_ : tendsto _ _ (𝓝 1)) one_ne_zero, + rw (by simp : (1 : ℂ) = real.exp (-0)), + exact (continuous_of_real.comp (real.continuous_exp.comp continuous_neg)).continuous_within_at } end lemma differentiable_at_Gamma_aux (s : ℂ) (n : ℕ) (h1 : (1 - s.re) < n ) (h2 : ∀ m : ℕ, s ≠ -m) : differentiable_at ℂ (Gamma_aux n) s := begin induction n with n hn generalizing s, - { refine (has_deriv_at_Gamma_integral _).2.differentiable_at, + { refine (has_deriv_at_Gamma_integral _).differentiable_at, rw nat.cast_zero at h1, linarith }, { dsimp only [Gamma_aux], specialize hn (s + 1), @@ -560,10 +445,10 @@ begin apply Gamma_eq_Gamma_aux, linarith, end -end complex - end Gamma_has_deriv +end complex + namespace real /-- The `Γ` function (of a real variable `s`). -/ diff --git a/src/analysis/special_functions/gamma/beta.lean b/src/analysis/special_functions/gamma/beta.lean index b396e1836867b..a30c97499cc4f 100644 --- a/src/analysis/special_functions/gamma/beta.lean +++ b/src/analysis/special_functions/gamma/beta.lean @@ -27,6 +27,7 @@ refined properties of the Gamma function using these relations. `n ↦ n ^ s * n! / (s * (s + 1) * ... * (s + n))` is `Γ(s)`. * `complex.Gamma_mul_Gamma_one_sub`: Euler's reflection formula `Gamma s * Gamma (1 - s) = π / sin π s`. +* `complex.differentiable_one_div_Gamma`: the function `1 / Γ(s)` is differentiable everywhere. * `real.Gamma_ne_zero`, `real.Gamma_seq_tendsto_Gamma`, `real.Gamma_mul_Gamma_one_sub`: real versions of the above results. -/ @@ -507,3 +508,47 @@ end end real end gamma_reflection + +section inv_gamma +open_locale real + +namespace complex +/-! ## The reciprocal Gamma function + +We show that the reciprocal Gamma function `1 / Γ(s)` is entire. These lemmas show that (in this +case at least) mathlib's conventions for division by zero do actually give a mathematically useful +answer! (These results are useful in the theory of zeta and L-functions.) -/ + +/-- A reformulation of the Gamma recurrence relation which is true for `s = 0` as well. -/ +lemma one_div_Gamma_eq_self_mul_one_div_Gamma_add_one (s : ℂ) : + (Gamma s)⁻¹ = s * (Gamma (s + 1))⁻¹ := +begin + rcases ne_or_eq s 0 with h | rfl, + { rw [Gamma_add_one s h, mul_inv, mul_inv_cancel_left₀ h] }, + { rw [zero_add, Gamma_zero, inv_zero, zero_mul] } +end + +/-- The reciprocal of the Gamma function is differentiable everywhere (including the points where +Gamma itself is not). -/ +lemma differentiable_one_div_Gamma : differentiable ℂ (λ s : ℂ, (Gamma s)⁻¹) := +begin + suffices : ∀ (n : ℕ), ∀ (s : ℂ) (hs : -s.re < n), differentiable_at ℂ (λ u : ℂ, (Gamma u)⁻¹) s, + from λ s, let ⟨n, h⟩ := exists_nat_gt (-s.re) in this n s h, + intro n, + induction n with m hm, + { intros s hs, + rw [nat.cast_zero, neg_lt_zero] at hs, + suffices : ∀ (m : ℕ), s ≠ -↑m, from (differentiable_at_Gamma _ this).inv (Gamma_ne_zero this), + contrapose! hs, + rcases hs with ⟨m, rfl⟩, + simpa only [neg_re, ←of_real_nat_cast, of_real_re, neg_nonpos] using nat.cast_nonneg m }, + { intros s hs, + rw funext one_div_Gamma_eq_self_mul_one_div_Gamma_add_one, + specialize hm (s + 1) (by rwa [add_re, one_re, neg_add', sub_lt_iff_lt_add, ←nat.cast_succ]), + refine differentiable_at_id.mul (hm.comp s _), + exact differentiable_at_id.add (differentiable_at_const _) } +end + +end complex + +end inv_gamma From f8c79b0a623404854a2902b836eac32156fd7712 Mon Sep 17 00:00:00 2001 From: Giovanni Mascellani Date: Tue, 16 May 2023 19:40:58 +0000 Subject: [PATCH 1018/1291] doc(linear_algebra/std_basis): add a missing backtick (#19025) Co-authored-by: Giovanni Mascellani --- src/linear_algebra/std_basis.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/linear_algebra/std_basis.lean b/src/linear_algebra/std_basis.lean index 0cc92187283bb..f1af076104f9f 100644 --- a/src/linear_algebra/std_basis.lean +++ b/src/linear_algebra/std_basis.lean @@ -14,7 +14,7 @@ import linear_algebra.pi > Any changes to this file require a corresponding PR to mathlib4. This file defines the standard basis `pi.basis (s : ∀ j, basis (ι j) R (M j))`, -which is the `Σ j, ι j`-indexed basis of Π j, M j`. The basis vectors are given by +which is the `Σ j, ι j`-indexed basis of `Π j, M j`. The basis vectors are given by `pi.basis s ⟨j, i⟩ j' = linear_map.std_basis R M j' (s j) i = if j = j' then s i else 0`. The standard basis on `R^η`, i.e. `η → R` is called `pi.basis_fun`. From b1859b6d4636fdbb78c5d5cefd24530653cfd3eb Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 16 May 2023 22:14:44 +0000 Subject: [PATCH 1019/1291] =?UTF-8?q?feat(measure=5Ftheory/measure/hausdor?= =?UTF-8?q?ff):=20`(=CE=BCH[1]=20:=20measure=20=E2=84=9D)=20=3D=20volume`?= =?UTF-8?q?=20(#18982)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit And similarly for `(μH[2] : measure ℝ × ℝ) = volume`. This addresses the TODO comment in the docstring. The `hausdorff_measure_pi_real` proof has been moved to the bottom of the file so that it can be kept next to the new results. --- src/measure_theory/measure/hausdorff.lean | 242 ++++++++++++---------- src/topology/metric_space/isometry.lean | 17 +- 2 files changed, 151 insertions(+), 108 deletions(-) diff --git a/src/measure_theory/measure/hausdorff.lean b/src/measure_theory/measure/hausdorff.lean index cf5ef8418662e..aa1cf0a4df7da 100644 --- a/src/measure_theory/measure/hausdorff.lean +++ b/src/measure_theory/measure/hausdorff.lean @@ -96,11 +96,6 @@ sources only allow coverings by balls and use `r ^ d` instead of `(diam s) ^ d`. construction lead to different Hausdorff measures, they lead to the same notion of the Hausdorff dimension. -## TODO - -* prove that `1`-dimensional Hausdorff measure on `ℝ` equals `volume`; -* prove a similar statement for `ℝ × ℝ`. - ## References * [Herbert Federer, Geometric Measure Theory, Chapter 2.10][Federer1996] @@ -693,108 +688,6 @@ end end measure -open_locale measure_theory -open measure - -/-! -### Hausdorff measure and Lebesgue measure --/ - -/-- In the space `ι → ℝ`, Hausdorff measure coincides exactly with Lebesgue measure. -/ -@[simp] theorem hausdorff_measure_pi_real {ι : Type*} [fintype ι] : - (μH[fintype.card ι] : measure (ι → ℝ)) = volume := -begin - classical, - -- it suffices to check that the two measures coincide on products of rational intervals - refine (pi_eq_generate_from (λ i, real.borel_eq_generate_from_Ioo_rat.symm) - (λ i, real.is_pi_system_Ioo_rat) (λ i, real.finite_spanning_sets_in_Ioo_rat _) - _).symm, - simp only [mem_Union, mem_singleton_iff], - -- fix such a product `s` of rational intervals, of the form `Π (a i, b i)`. - intros s hs, - choose a b H using hs, - obtain rfl : s = λ i, Ioo (a i) (b i), from funext (λ i, (H i).2), replace H := λ i, (H i).1, - apply le_antisymm _, - -- first check that `volume s ≤ μH s` - { have Hle : volume ≤ (μH[fintype.card ι] : measure (ι → ℝ)), - { refine le_hausdorff_measure _ _ ∞ ennreal.coe_lt_top (λ s _, _), - rw [ennreal.rpow_nat_cast], - exact real.volume_pi_le_diam_pow s }, - rw [← volume_pi_pi (λ i, Ioo (a i : ℝ) (b i))], - exact measure.le_iff'.1 Hle _ }, - /- For the other inequality `μH s ≤ volume s`, we use a covering of `s` by sets of small diameter - `1/n`, namely cubes with left-most point of the form `a i + f i / n` with `f i` ranging between - `0` and `⌈(b i - a i) * n⌉`. Their number is asymptotic to `n^d * Π (b i - a i)`. -/ - have I : ∀ i, 0 ≤ (b i : ℝ) - a i := λ i, by simpa only [sub_nonneg, rat.cast_le] using (H i).le, - let γ := λ (n : ℕ), (Π (i : ι), fin ⌈((b i : ℝ) - a i) * n⌉₊), - let t : Π (n : ℕ), γ n → set (ι → ℝ) := - λ n f, set.pi univ (λ i, Icc (a i + f i / n) (a i + (f i + 1) / n)), - have A : tendsto (λ (n : ℕ), 1/(n : ℝ≥0∞)) at_top (𝓝 0), - by simp only [one_div, ennreal.tendsto_inv_nat_nhds_zero], - have B : ∀ᶠ n in at_top, ∀ (i : γ n), diam (t n i) ≤ 1 / n, - { apply eventually_at_top.2 ⟨1, λ n hn, _⟩, - assume f, - apply diam_pi_le_of_le (λ b, _), - simp only [real.ediam_Icc, add_div, ennreal.of_real_div_of_pos (nat.cast_pos.mpr hn), le_refl, - add_sub_add_left_eq_sub, add_sub_cancel', ennreal.of_real_one, ennreal.of_real_coe_nat] }, - have C : ∀ᶠ n in at_top, set.pi univ (λ (i : ι), Ioo (a i : ℝ) (b i)) ⊆ ⋃ (i : γ n), t n i, - { apply eventually_at_top.2 ⟨1, λ n hn, _⟩, - have npos : (0 : ℝ) < n := nat.cast_pos.2 hn, - assume x hx, - simp only [mem_Ioo, mem_univ_pi] at hx, - simp only [mem_Union, mem_Ioo, mem_univ_pi, coe_coe], - let f : γ n := λ i, ⟨⌊(x i - a i) * n⌋₊, - begin - apply nat.floor_lt_ceil_of_lt_of_pos, - { refine (mul_lt_mul_right npos).2 _, - simp only [(hx i).right, sub_lt_sub_iff_right] }, - { refine mul_pos _ npos, - simpa only [rat.cast_lt, sub_pos] using H i } - end⟩, - refine ⟨f, λ i, ⟨_, _⟩⟩, - { calc (a i : ℝ) + ⌊(x i - a i) * n⌋₊ / n - ≤ (a i : ℝ) + ((x i - a i) * n) / n : - begin - refine add_le_add le_rfl ((div_le_div_right npos).2 _), - exact nat.floor_le (mul_nonneg (sub_nonneg.2 (hx i).1.le) npos.le), - end - ... = x i : by field_simp [npos.ne'] }, - { calc x i - = (a i : ℝ) + ((x i - a i) * n) / n : by field_simp [npos.ne'] - ... ≤ (a i : ℝ) + (⌊(x i - a i) * n⌋₊ + 1) / n : - add_le_add le_rfl ((div_le_div_right npos).2 (nat.lt_floor_add_one _).le) } }, - calc μH[fintype.card ι] (set.pi univ (λ (i : ι), Ioo (a i : ℝ) (b i))) - ≤ liminf (λ (n : ℕ), ∑ (i : γ n), diam (t n i) ^ ↑(fintype.card ι)) at_top : - hausdorff_measure_le_liminf_sum _ (set.pi univ (λ i, Ioo (a i : ℝ) (b i))) - (λ (n : ℕ), 1/(n : ℝ≥0∞)) A t B C - ... ≤ liminf (λ (n : ℕ), ∑ (i : γ n), (1/n) ^ (fintype.card ι)) at_top : - begin - refine liminf_le_liminf _ (by is_bounded_default), - filter_upwards [B] with _ hn, - apply finset.sum_le_sum (λ i _, _), - rw ennreal.rpow_nat_cast, - exact pow_le_pow_of_le_left' (hn i) _, - end - ... = liminf (λ (n : ℕ), ∏ (i : ι), (⌈((b i : ℝ) - a i) * n⌉₊ : ℝ≥0∞) / n) at_top : - begin - simp only [finset.card_univ, nat.cast_prod, one_mul, fintype.card_fin, - finset.sum_const, nsmul_eq_mul, fintype.card_pi, div_eq_mul_inv, finset.prod_mul_distrib, - finset.prod_const] - end - ... = ∏ (i : ι), volume (Ioo (a i : ℝ) (b i)) : - begin - simp only [real.volume_Ioo], - apply tendsto.liminf_eq, - refine ennreal.tendsto_finset_prod_of_ne_top _ (λ i hi, _) (λ i hi, _), - { apply tendsto.congr' _ ((ennreal.continuous_of_real.tendsto _).comp - ((tendsto_nat_ceil_mul_div_at_top (I i)).comp tendsto_coe_nat_at_top_at_top)), - apply eventually_at_top.2 ⟨1, λ n hn, _⟩, - simp only [ennreal.of_real_div_of_pos (nat.cast_pos.mpr hn), comp_app, - ennreal.of_real_coe_nat] }, - { simp only [ennreal.of_real_ne_top, ne.def, not_false_iff] } - end -end - end measure_theory /-! @@ -964,4 +857,139 @@ e.isometry.hausdorff_measure_image (or.inr e.surjective) s μH[d] (e ⁻¹' s) = μH[d] s := by rw [← e.image_symm, e.symm.hausdorff_measure_image] +@[simp] lemma map_hausdorff_measure (e : X ≃ᵢ Y) (d : ℝ) : measure.map e μH[d] = μH[d] := +by rw [e.isometry.map_hausdorff_measure (or.inr e.surjective), e.surjective.range_eq, restrict_univ] + +lemma measure_preserving_hausdorff_measure (e : X ≃ᵢ Y) (d : ℝ) : + measure_preserving e μH[d] μH[d] := +⟨e.continuous.measurable, map_hausdorff_measure _ _⟩ + end isometry_equiv + +namespace measure_theory + +/-! +### Hausdorff measure and Lebesgue measure +-/ + +/-- In the space `ι → ℝ`, the Hausdorff measure coincides exactly with the Lebesgue measure. -/ +@[simp] theorem hausdorff_measure_pi_real {ι : Type*} [fintype ι] : + (μH[fintype.card ι] : measure (ι → ℝ)) = volume := +begin + classical, + -- it suffices to check that the two measures coincide on products of rational intervals + refine (pi_eq_generate_from (λ i, real.borel_eq_generate_from_Ioo_rat.symm) + (λ i, real.is_pi_system_Ioo_rat) (λ i, real.finite_spanning_sets_in_Ioo_rat _) + _).symm, + simp only [mem_Union, mem_singleton_iff], + -- fix such a product `s` of rational intervals, of the form `Π (a i, b i)`. + intros s hs, + choose a b H using hs, + obtain rfl : s = λ i, Ioo (a i) (b i), from funext (λ i, (H i).2), replace H := λ i, (H i).1, + apply le_antisymm _, + -- first check that `volume s ≤ μH s` + { have Hle : volume ≤ (μH[fintype.card ι] : measure (ι → ℝ)), + { refine le_hausdorff_measure _ _ ∞ ennreal.coe_lt_top (λ s _, _), + rw [ennreal.rpow_nat_cast], + exact real.volume_pi_le_diam_pow s }, + rw [← volume_pi_pi (λ i, Ioo (a i : ℝ) (b i))], + exact measure.le_iff'.1 Hle _ }, + /- For the other inequality `μH s ≤ volume s`, we use a covering of `s` by sets of small diameter + `1/n`, namely cubes with left-most point of the form `a i + f i / n` with `f i` ranging between + `0` and `⌈(b i - a i) * n⌉`. Their number is asymptotic to `n^d * Π (b i - a i)`. -/ + have I : ∀ i, 0 ≤ (b i : ℝ) - a i := λ i, by simpa only [sub_nonneg, rat.cast_le] using (H i).le, + let γ := λ (n : ℕ), (Π (i : ι), fin ⌈((b i : ℝ) - a i) * n⌉₊), + let t : Π (n : ℕ), γ n → set (ι → ℝ) := + λ n f, set.pi univ (λ i, Icc (a i + f i / n) (a i + (f i + 1) / n)), + have A : tendsto (λ (n : ℕ), 1/(n : ℝ≥0∞)) at_top (𝓝 0), + by simp only [one_div, ennreal.tendsto_inv_nat_nhds_zero], + have B : ∀ᶠ n in at_top, ∀ (i : γ n), diam (t n i) ≤ 1 / n, + { apply eventually_at_top.2 ⟨1, λ n hn, _⟩, + assume f, + apply diam_pi_le_of_le (λ b, _), + simp only [real.ediam_Icc, add_div, ennreal.of_real_div_of_pos (nat.cast_pos.mpr hn), le_refl, + add_sub_add_left_eq_sub, add_sub_cancel', ennreal.of_real_one, ennreal.of_real_coe_nat] }, + have C : ∀ᶠ n in at_top, set.pi univ (λ (i : ι), Ioo (a i : ℝ) (b i)) ⊆ ⋃ (i : γ n), t n i, + { apply eventually_at_top.2 ⟨1, λ n hn, _⟩, + have npos : (0 : ℝ) < n := nat.cast_pos.2 hn, + assume x hx, + simp only [mem_Ioo, mem_univ_pi] at hx, + simp only [mem_Union, mem_Ioo, mem_univ_pi, coe_coe], + let f : γ n := λ i, ⟨⌊(x i - a i) * n⌋₊, + begin + apply nat.floor_lt_ceil_of_lt_of_pos, + { refine (mul_lt_mul_right npos).2 _, + simp only [(hx i).right, sub_lt_sub_iff_right] }, + { refine mul_pos _ npos, + simpa only [rat.cast_lt, sub_pos] using H i } + end⟩, + refine ⟨f, λ i, ⟨_, _⟩⟩, + { calc (a i : ℝ) + ⌊(x i - a i) * n⌋₊ / n + ≤ (a i : ℝ) + ((x i - a i) * n) / n : + begin + refine add_le_add le_rfl ((div_le_div_right npos).2 _), + exact nat.floor_le (mul_nonneg (sub_nonneg.2 (hx i).1.le) npos.le), + end + ... = x i : by field_simp [npos.ne'] }, + { calc x i + = (a i : ℝ) + ((x i - a i) * n) / n : by field_simp [npos.ne'] + ... ≤ (a i : ℝ) + (⌊(x i - a i) * n⌋₊ + 1) / n : + add_le_add le_rfl ((div_le_div_right npos).2 (nat.lt_floor_add_one _).le) } }, + calc μH[fintype.card ι] (set.pi univ (λ (i : ι), Ioo (a i : ℝ) (b i))) + ≤ liminf (λ (n : ℕ), ∑ (i : γ n), diam (t n i) ^ ↑(fintype.card ι)) at_top : + hausdorff_measure_le_liminf_sum _ (set.pi univ (λ i, Ioo (a i : ℝ) (b i))) + (λ (n : ℕ), 1/(n : ℝ≥0∞)) A t B C + ... ≤ liminf (λ (n : ℕ), ∑ (i : γ n), (1/n) ^ (fintype.card ι)) at_top : + begin + refine liminf_le_liminf _ (by is_bounded_default), + filter_upwards [B] with _ hn, + apply finset.sum_le_sum (λ i _, _), + rw ennreal.rpow_nat_cast, + exact pow_le_pow_of_le_left' (hn i) _, + end + ... = liminf (λ (n : ℕ), ∏ (i : ι), (⌈((b i : ℝ) - a i) * n⌉₊ : ℝ≥0∞) / n) at_top : + begin + simp only [finset.card_univ, nat.cast_prod, one_mul, fintype.card_fin, + finset.sum_const, nsmul_eq_mul, fintype.card_pi, div_eq_mul_inv, finset.prod_mul_distrib, + finset.prod_const] + end + ... = ∏ (i : ι), volume (Ioo (a i : ℝ) (b i)) : + begin + simp only [real.volume_Ioo], + apply tendsto.liminf_eq, + refine ennreal.tendsto_finset_prod_of_ne_top _ (λ i hi, _) (λ i hi, _), + { apply tendsto.congr' _ ((ennreal.continuous_of_real.tendsto _).comp + ((tendsto_nat_ceil_mul_div_at_top (I i)).comp tendsto_coe_nat_at_top_at_top)), + apply eventually_at_top.2 ⟨1, λ n hn, _⟩, + simp only [ennreal.of_real_div_of_pos (nat.cast_pos.mpr hn), comp_app, + ennreal.of_real_coe_nat] }, + { simp only [ennreal.of_real_ne_top, ne.def, not_false_iff] } + end +end + +variables (ι X) + +theorem hausdorff_measure_measure_preserving_fun_unique [unique ι] + [topological_space.second_countable_topology X] (d : ℝ) : + measure_preserving (measurable_equiv.fun_unique ι X) μH[d] μH[d] := +(isometry_equiv.fun_unique ι X).measure_preserving_hausdorff_measure _ + +theorem hausdorff_measure_measure_preserving_pi_fin_two (α : fin 2 → Type*) + [Π i, measurable_space (α i)] [Π i, emetric_space (α i)] [Π i, borel_space (α i)] + [Π i, topological_space.second_countable_topology (α i)] (d : ℝ) : + measure_preserving (measurable_equiv.pi_fin_two α) μH[d] μH[d] := +(isometry_equiv.pi_fin_two α).measure_preserving_hausdorff_measure _ + +/-- In the space `ℝ`, the Hausdorff measure coincides exactly with the Lebesgue measure. -/ +@[simp] theorem hausdorff_measure_real : (μH[1] : measure ℝ) = volume := +by rw [←(volume_preserving_fun_unique unit ℝ).map_eq, + ←(hausdorff_measure_measure_preserving_fun_unique unit ℝ 1).map_eq, + ←hausdorff_measure_pi_real, fintype.card_unit, nat.cast_one] + +/-- In the space `ℝ × ℝ`, the Hausdorff measure coincides exactly with the Lebesgue measure. -/ +@[simp] theorem hausdorff_measure_prod_real : (μH[2] : measure (ℝ × ℝ)) = volume := +by rw [←(volume_preserving_pi_fin_two (λ i, ℝ)).map_eq, + ←(hausdorff_measure_measure_preserving_pi_fin_two (λ i, ℝ) _).map_eq, + ←hausdorff_measure_pi_real, fintype.card_fin, nat.cast_two] + +end measure_theory diff --git a/src/topology/metric_space/isometry.lean b/src/topology/metric_space/isometry.lean index ef29d1dcb3200..0836c9fe2fb95 100644 --- a/src/topology/metric_space/isometry.lean +++ b/src/topology/metric_space/isometry.lean @@ -23,7 +23,7 @@ theory for `pseudo_metric_space` and we specialize to `metric_space` when needed noncomputable theory universes u v w -variables {α : Type u} {β : Type v} {γ : Type w} +variables {ι : Type*} {α : Type u} {β : Type v} {γ : Type w} open function set open_locale topology ennreal @@ -431,6 +431,21 @@ complete_space_of_is_complete_univ $ is_complete_of_complete_image e.isometry.un lemma complete_space_iff (e : α ≃ᵢ β) : complete_space α ↔ complete_space β := by { split; introI H, exacts [e.symm.complete_space, e.complete_space] } +variables (ι α) + +/-- `equiv.fun_unique` as an `isometry_equiv`. -/ +@[simps] +def fun_unique [unique ι] [fintype ι] : (ι → α) ≃ᵢ α := +{ to_equiv := equiv.fun_unique ι α, + isometry_to_fun := λ x hx, by simp [edist_pi_def, finset.univ_unique, finset.sup_singleton] } + +/-- `pi_fin_two_equiv` as an `isometry_equiv`. -/ +@[simps] +def pi_fin_two (α : fin 2 → Type*) [Π i, pseudo_emetric_space (α i)] : + (Π i, α i) ≃ᵢ α 0 × α 1 := +{ to_equiv := pi_fin_two_equiv α, + isometry_to_fun := λ x hx, by simp [edist_pi_def, fin.univ_succ, prod.edist_eq] } + end pseudo_emetric_space section pseudo_metric_space From 2ed2c6310e6f1c5562bdf6bfbda55ebbf6891abe Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Wed, 17 May 2023 00:23:15 +0000 Subject: [PATCH 1020/1291] chore(*): add mathlib4 synchronization comments (#19024) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.module.localized_module` * `analysis.box_integral.partition.tagged` * `analysis.calculus.tangent_cone` * `analysis.complex.basic` * `analysis.complex.re_im_topology` * `analysis.normed.group.add_circle` * `analysis.special_functions.exp` * `analysis.special_functions.log.basic` * `category_theory.cofiltered_system` * `category_theory.endofunctor.algebra` * `category_theory.sites.dense_subsite` * `category_theory.sites.induced_topology` * `category_theory.sites.subsheaf` * `combinatorics.hall.basic` * `combinatorics.simple_graph.ends.defs` * `data.polynomial.laurent` * `group_theory.nielsen_schreier` * `order.category.Frm` * `topology.category.Locale` * `topology.instances.add_circle` --- src/algebra/module/localized_module.lean | 3 +++ src/analysis/box_integral/partition/tagged.lean | 3 +++ src/analysis/calculus/tangent_cone.lean | 3 +++ src/analysis/complex/basic.lean | 3 +++ src/analysis/complex/re_im_topology.lean | 3 +++ src/analysis/normed/group/add_circle.lean | 3 +++ src/analysis/special_functions/exp.lean | 3 +++ src/analysis/special_functions/log/basic.lean | 3 +++ src/category_theory/cofiltered_system.lean | 3 +++ src/category_theory/endofunctor/algebra.lean | 3 +++ src/category_theory/sites/dense_subsite.lean | 3 +++ src/category_theory/sites/induced_topology.lean | 3 +++ src/category_theory/sites/subsheaf.lean | 3 +++ src/combinatorics/hall/basic.lean | 3 +++ src/combinatorics/simple_graph/ends/defs.lean | 3 +++ src/data/polynomial/laurent.lean | 3 +++ src/group_theory/nielsen_schreier.lean | 3 +++ src/order/category/Frm.lean | 3 +++ src/topology/category/Locale.lean | 3 +++ src/topology/instances/add_circle.lean | 3 +++ 20 files changed, 60 insertions(+) diff --git a/src/algebra/module/localized_module.lean b/src/algebra/module/localized_module.lean index 8941bab7c10a4..7df4bd5427c3b 100644 --- a/src/algebra/module/localized_module.lean +++ b/src/algebra/module/localized_module.lean @@ -11,6 +11,9 @@ import algebra.algebra.restrict_scalars /-! # Localized Module +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given a commutative ring `R`, a multiplicative subset `S ⊆ R` and an `R`-module `M`, we can localize `M` by `S`. This gives us a `localization S`-module. diff --git a/src/analysis/box_integral/partition/tagged.lean b/src/analysis/box_integral/partition/tagged.lean index 8bd1c84db7902..ea6c3b12ebe9b 100644 --- a/src/analysis/box_integral/partition/tagged.lean +++ b/src/analysis/box_integral/partition/tagged.lean @@ -8,6 +8,9 @@ import analysis.box_integral.partition.basic /-! # Tagged partitions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A tagged (pre)partition is a (pre)partition `π` enriched with a tagged point for each box of ‵π`. For simplicity we require that the function `box_integral.tagged_prepartition.tag` is defined on all boxes `J : box ι` but use its values only on boxes of the partition. Given `π : diff --git a/src/analysis/calculus/tangent_cone.lean b/src/analysis/calculus/tangent_cone.lean index 10009a260f271..3759c2c22d93a 100644 --- a/src/analysis/calculus/tangent_cone.lean +++ b/src/analysis/calculus/tangent_cone.lean @@ -10,6 +10,9 @@ import analysis.specific_limits.basic /-! # Tangent cone +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we define two predicates `unique_diff_within_at 𝕜 s x` and `unique_diff_on 𝕜 s` ensuring that, if a function has two derivatives, then they have to coincide. As a direct definition of this fact (quantifying on all target types and all functions) would depend on diff --git a/src/analysis/complex/basic.lean b/src/analysis/complex/basic.lean index a475d7eca7f07..9b8097b45dcf2 100644 --- a/src/analysis/complex/basic.lean +++ b/src/analysis/complex/basic.lean @@ -12,6 +12,9 @@ import topology.instances.real_vector_space /-! # Normed space structure on `ℂ`. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file gathers basic facts on complex numbers of an analytic nature. ## Main results diff --git a/src/analysis/complex/re_im_topology.lean b/src/analysis/complex/re_im_topology.lean index 5410395907c29..e86d1e2b077d3 100644 --- a/src/analysis/complex/re_im_topology.lean +++ b/src/analysis/complex/re_im_topology.lean @@ -9,6 +9,9 @@ import topology.fiber_bundle.is_homeomorphic_trivial_bundle /-! # Closure, interior, and frontier of preimages under `re` and `im` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this fact we use the fact that `ℂ` is naturally homeomorphic to `ℝ × ℝ` to deduce some topological properties of `complex.re` and `complex.im`. diff --git a/src/analysis/normed/group/add_circle.lean b/src/analysis/normed/group/add_circle.lean index 192c1f5c0a213..d62aa04e4c021 100644 --- a/src/analysis/normed/group/add_circle.lean +++ b/src/analysis/normed/group/add_circle.lean @@ -9,6 +9,9 @@ import topology.instances.add_circle /-! # The additive circle as a normed group +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define the normed group structure on `add_circle p`, for `p : ℝ`. For example if `p = 1` then: `‖(x : add_circle 1)‖ = |x - round x|` for any `x : ℝ` (see `unit_add_circle.norm_eq`). diff --git a/src/analysis/special_functions/exp.lean b/src/analysis/special_functions/exp.lean index b1b377aef9ff6..02e03f130d6b6 100644 --- a/src/analysis/special_functions/exp.lean +++ b/src/analysis/special_functions/exp.lean @@ -10,6 +10,9 @@ import analysis.specific_limits.normed /-! # Complex and real exponential +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove continuity of `complex.exp` and `real.exp`. We also prove a few facts about limits of `real.exp` at infinity. diff --git a/src/analysis/special_functions/log/basic.lean b/src/analysis/special_functions/log/basic.lean index d70fb1ebd8237..60b4961efe531 100644 --- a/src/analysis/special_functions/log/basic.lean +++ b/src/analysis/special_functions/log/basic.lean @@ -9,6 +9,9 @@ import data.nat.factorization.basic /-! # Real logarithm +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define `real.log` to be the logarithm of a real number. As usual, we extend it from its domain `(0, +∞)` to a globally defined function. We choose to do it so that `log 0 = 0` and `log (-x) = log x`. diff --git a/src/category_theory/cofiltered_system.lean b/src/category_theory/cofiltered_system.lean index 73df182bb6ee8..1fa9fddb32a8d 100644 --- a/src/category_theory/cofiltered_system.lean +++ b/src/category_theory/cofiltered_system.lean @@ -10,6 +10,9 @@ import topology.category.Top.limits.konig /-! # Cofiltered systems +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file deals with properties of cofiltered (and inverse) systems. ## Main definitions diff --git a/src/category_theory/endofunctor/algebra.lean b/src/category_theory/endofunctor/algebra.lean index 436f6db1bc616..42a453ec5d005 100644 --- a/src/category_theory/endofunctor/algebra.lean +++ b/src/category_theory/endofunctor/algebra.lean @@ -10,6 +10,9 @@ import category_theory.limits.shapes.terminal # Algebras of endofunctors +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines (co)algebras of an endofunctor, and provides the category instance for them. It also defines the forgetful functor from the category of (co)algebras. It is shown that the structure map of the initial algebra of an endofunctor is an isomorphism. Furthermore, it is shown diff --git a/src/category_theory/sites/dense_subsite.lean b/src/category_theory/sites/dense_subsite.lean index d17ad4ce184ca..8a2498b39b159 100644 --- a/src/category_theory/sites/dense_subsite.lean +++ b/src/category_theory/sites/dense_subsite.lean @@ -10,6 +10,9 @@ import category_theory.adjunction.fully_faithful /-! # Dense subsites +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define `cover_dense` functors into sites as functors such that there exists a covering sieve that factors through images of the functor for each object in `D`. diff --git a/src/category_theory/sites/induced_topology.lean b/src/category_theory/sites/induced_topology.lean index 7f0cd626e339a..8c54a1624e131 100644 --- a/src/category_theory/sites/induced_topology.lean +++ b/src/category_theory/sites/induced_topology.lean @@ -8,6 +8,9 @@ import category_theory.sites.dense_subsite /-! # Induced Topology +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We say that a functor `G : C ⥤ (D, K)` is locally dense if for each covering sieve `T` in `D` of some `X : C`, `T ∩ mor(C)` generates a covering sieve of `X` in `D`. A locally dense fully faithful functor then induces a topology on `C` via `{ T ∩ mor(C) | T ∈ K }`. Note that this is equal to diff --git a/src/category_theory/sites/subsheaf.lean b/src/category_theory/sites/subsheaf.lean index 1d5cf76d944ff..b0bd284af7278 100644 --- a/src/category_theory/sites/subsheaf.lean +++ b/src/category_theory/sites/subsheaf.lean @@ -11,6 +11,9 @@ import category_theory.sites.sheafification # Subsheaf of types +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define the sub(pre)sheaf of a type valued presheaf. ## Main results diff --git a/src/combinatorics/hall/basic.lean b/src/combinatorics/hall/basic.lean index 93ab54529b959..b150a76c13cf9 100644 --- a/src/combinatorics/hall/basic.lean +++ b/src/combinatorics/hall/basic.lean @@ -10,6 +10,9 @@ import data.rel /-! # Hall's Marriage Theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given a list of finite subsets $X_1, X_2, \dots, X_n$ of some given set $S$, P. Hall in [Hall1935] gave a necessary and sufficient condition for there to be a list of distinct elements $x_1, x_2, \dots, x_n$ with diff --git a/src/combinatorics/simple_graph/ends/defs.lean b/src/combinatorics/simple_graph/ends/defs.lean index 3f5c634f5f3d1..87889d5d0610f 100644 --- a/src/combinatorics/simple_graph/ends/defs.lean +++ b/src/combinatorics/simple_graph/ends/defs.lean @@ -10,6 +10,9 @@ import data.set_like.basic /-! # Ends +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains a definition of the ends of a simple graph, as sections of the inverse system assigning, to each finite set of vertices, the connected components of its complement. -/ diff --git a/src/data/polynomial/laurent.lean b/src/data/polynomial/laurent.lean index bcdfada8689be..d89c24a6a75f0 100644 --- a/src/data/polynomial/laurent.lean +++ b/src/data/polynomial/laurent.lean @@ -9,6 +9,9 @@ import ring_theory.localization.basic /-! # Laurent polynomials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We introduce Laurent polynomials over a semiring `R`. Mathematically, they are expressions of the form $$ diff --git a/src/group_theory/nielsen_schreier.lean b/src/group_theory/nielsen_schreier.lean index a54d8f0d7a83b..03c68bbb4ad18 100644 --- a/src/group_theory/nielsen_schreier.lean +++ b/src/group_theory/nielsen_schreier.lean @@ -10,6 +10,9 @@ import group_theory.is_free_group /-! # The Nielsen-Schreier theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves that a subgroup of a free group is itself free. ## Main result diff --git a/src/order/category/Frm.lean b/src/order/category/Frm.lean index 61e67e1efa97e..6052f12eb8b4f 100644 --- a/src/order/category/Frm.lean +++ b/src/order/category/Frm.lean @@ -11,6 +11,9 @@ import topology.sets.opens /-! # The category of frames +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines `Frm`, the category of frames. ## References diff --git a/src/topology/category/Locale.lean b/src/topology/category/Locale.lean index 795f926e4ebc8..2bfa9589c3744 100644 --- a/src/topology/category/Locale.lean +++ b/src/topology/category/Locale.lean @@ -8,6 +8,9 @@ import order.category.Frm /-! # The category of locales +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines `Locale`, the category of locales. This is the opposite of the category of frames. -/ diff --git a/src/topology/instances/add_circle.lean b/src/topology/instances/add_circle.lean index e84d83778331f..88ac997cbf3eb 100644 --- a/src/topology/instances/add_circle.lean +++ b/src/topology/instances/add_circle.lean @@ -14,6 +14,9 @@ import topology.instances.real /-! # The additive circle +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define the additive circle `add_circle p` as the quotient `𝕜 ⧸ (ℤ ∙ p)` for some period `p : 𝕜`. See also `circle` and `real.angle`. For the normed group structure on `add_circle`, see From 20d5763051978e9bc6428578ed070445df6a18b3 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Wed, 17 May 2023 10:26:04 +0000 Subject: [PATCH 1021/1291] feat(probability/kernel/cond_cdf): conditional cumulative distribution function (#18988) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We define `cond_cdf ρ : α → stieltjes_function`, the conditional cdf of `ρ : measure (α × ℝ)`, and prove its main properties: it is monotone and right-continuous with limit 0 at -∞ and limit 1 at +∞; for all `x : ℝ`, `a ↦ cond_cdf ρ a x` is measurable; for all `x : ℝ` and measurable set `s`, it satisfies `∫⁻ a in s, ennreal.of_real (cond_cdf ρ a x) ∂ρ.fst = ρ (s ×ˢ Iic x)` --- .../constructions/borel_space/basic.lean | 17 + src/measure_theory/measure/stieltjes.lean | 68 ++ src/probability/kernel/cond_cdf.lean | 969 ++++++++++++++++++ 3 files changed, 1054 insertions(+) create mode 100644 src/probability/kernel/cond_cdf.lean diff --git a/src/measure_theory/constructions/borel_space/basic.lean b/src/measure_theory/constructions/borel_space/basic.lean index c52b04a8be156..9ca71671a64b3 100644 --- a/src/measure_theory/constructions/borel_space/basic.lean +++ b/src/measure_theory/constructions/borel_space/basic.lean @@ -148,6 +148,23 @@ end lemma borel_eq_generate_from_Ioi : borel α = generate_from (range Ioi) := @borel_eq_generate_from_Iio αᵒᵈ _ (by apply_instance : second_countable_topology α) _ _ +lemma borel_eq_generate_from_Iic : borel α = measurable_space.generate_from (range Iic) := +begin + rw borel_eq_generate_from_Ioi, + refine le_antisymm _ _, + { refine measurable_space.generate_from_le (λ t ht, _), + obtain ⟨u, rfl⟩ := ht, + rw ← compl_Iic, + exact (measurable_space.measurable_set_generate_from (mem_range.mpr ⟨u, rfl⟩)).compl, }, + { refine measurable_space.generate_from_le (λ t ht, _), + obtain ⟨u, rfl⟩ := ht, + rw ← compl_Ioi, + exact (measurable_space.measurable_set_generate_from (mem_range.mpr ⟨u, rfl⟩)).compl, }, +end + +lemma borel_eq_generate_from_Ici : borel α = measurable_space.generate_from (range Ici) := +@borel_eq_generate_from_Iic αᵒᵈ _ _ _ _ + end order_topology lemma borel_comap {f : α → β} {t : topological_space β} : diff --git a/src/measure_theory/measure/stieltjes.lean b/src/measure_theory/measure/stieltjes.lean index 07f34e0b845d4..7115704cf6288 100644 --- a/src/measure_theory/measure/stieltjes.lean +++ b/src/measure_theory/measure/stieltjes.lean @@ -29,6 +29,47 @@ section move_this open filter set open_locale topology +lemma infi_Ioi_eq_infi_rat_gt {f : ℝ → ℝ} (x : ℝ) (hf : bdd_below (f '' Ioi x)) + (hf_mono : monotone f) : + (⨅ r : Ioi x, f r) = ⨅ q : {q' : ℚ // x < q'}, f q := +begin + refine le_antisymm _ _, + { haveI : nonempty {r' : ℚ // x < ↑r'}, + { obtain ⟨r, hrx⟩ := exists_rat_gt x, + exact ⟨⟨r, hrx⟩⟩, }, + refine le_cinfi (λ r, _), + obtain ⟨y, hxy, hyr⟩ := exists_rat_btwn r.prop, + refine cinfi_set_le hf (hxy.trans _), + exact_mod_cast hyr, }, + { refine le_cinfi (λ q, _), + have hq := q.prop, + rw mem_Ioi at hq, + obtain ⟨y, hxy, hyq⟩ := exists_rat_btwn hq, + refine (cinfi_le _ _).trans _, + { exact ⟨y, hxy⟩, }, + { refine ⟨hf.some, λ z, _⟩, + rintros ⟨u, rfl⟩, + suffices hfu : f u ∈ f '' Ioi x, from hf.some_spec hfu, + exact ⟨u, u.prop, rfl⟩, }, + { refine hf_mono (le_trans _ hyq.le), + norm_cast, }, }, +end + +-- todo after the port: move to topology/algebra/order/left_right_lim +lemma right_lim_eq_of_tendsto {α β : Type*} [linear_order α] [topological_space β] + [hα : topological_space α] [h'α : order_topology α] [t2_space β] + {f : α → β} {a : α} {y : β} (h : 𝓝[>] a ≠ ⊥) (h' : tendsto f (𝓝[>] a) (𝓝 y)) : + function.right_lim f a = y := +@left_lim_eq_of_tendsto αᵒᵈ _ _ _ _ _ _ f a y h h' + +-- todo after the port: move to topology/algebra/order/left_right_lim +lemma right_lim_eq_Inf {α β : Type*} [linear_order α] [topological_space β] + [conditionally_complete_linear_order β] [order_topology β] {f : α → β} + (hf : monotone f) {x : α} + [topological_space α] [order_topology α] (h : 𝓝[>] x ≠ ⊥) : + function.right_lim f x = Inf (f '' (Ioi x)) := +right_lim_eq_of_tendsto h (hf.tendsto_nhds_within_Ioi x) + -- todo after the port: move to order/filter/at_top_bot lemma exists_seq_monotone_tendsto_at_top_at_top (α : Type*) [semilattice_sup α] [nonempty α] [(at_top : filter α).is_countably_generated] : @@ -172,6 +213,33 @@ lemma mono : monotone f := f.mono' lemma right_continuous (x : ℝ) : continuous_within_at f (Ici x) x := f.right_continuous' x +lemma right_lim_eq (f : stieltjes_function) (x : ℝ) : + function.right_lim f x = f x := +begin + rw [← f.mono.continuous_within_at_Ioi_iff_right_lim_eq, continuous_within_at_Ioi_iff_Ici], + exact f.right_continuous' x, +end + +lemma infi_Ioi_eq (f : stieltjes_function) (x : ℝ) : + (⨅ r : Ioi x, f r) = f x := +begin + suffices : function.right_lim f x = ⨅ r : Ioi x, f r, + { rw [← this, f.right_lim_eq], }, + rw [right_lim_eq_Inf f.mono, Inf_image'], + rw ← ne_bot_iff, + apply_instance, +end + +lemma infi_rat_gt_eq (f : stieltjes_function) (x : ℝ) : + (⨅ r : {r' : ℚ // x < r'}, f r) = f x := +begin + rw ← infi_Ioi_eq f x, + refine (infi_Ioi_eq_infi_rat_gt _ _ f.mono).symm, + refine ⟨f x, λ y, _⟩, + rintros ⟨y, hy_mem, rfl⟩, + exact f.mono (le_of_lt hy_mem), +end + /-- The identity of `ℝ` as a Stieltjes function, used to construct Lebesgue measure. -/ @[simps] protected def id : stieltjes_function := { to_fun := id, diff --git a/src/probability/kernel/cond_cdf.lean b/src/probability/kernel/cond_cdf.lean new file mode 100644 index 0000000000000..a98bbd76489a1 --- /dev/null +++ b/src/probability/kernel/cond_cdf.lean @@ -0,0 +1,969 @@ +/- +Copyright (c) 2023 Rémy Degenne. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Rémy Degenne +-/ +import measure_theory.measure.stieltjes +import probability.kernel.composition +import measure_theory.decomposition.radon_nikodym + +/-! +# Conditional cumulative distribution function + +Given `ρ : measure (α × ℝ)`, we define the conditional cumulative distribution function +(conditional cdf) of `ρ`. It is a function `cond_cdf ρ : α → ℝ → ℝ` such that if `ρ` is a finite +measure, then for all `a : α` `cond_cdf ρ a` is monotone and right-continuous with limit 0 at -∞ +and limit 1 at +∞, and such that for all `x : ℝ`, `a ↦ cond_cdf ρ a x` is measurable. For all +`x : ℝ` and measurable set `s`, that function satisfies +`∫⁻ a in s, ennreal.of_real (cond_cdf ρ a x) ∂ρ.fst = ρ (s ×ˢ Iic x)`. + +## Main definitions + +* `probability_theory.cond_cdf ρ : α → stieltjes_function`: the conditional cdf of + `ρ : measure (α × ℝ)`. A `stieltjes_function` is a function `ℝ → ℝ` which is monotone and + right-continuous. + +## Main statements + +* `probability_theory.set_lintegral_cond_cdf`: for all `a : α` and `x : ℝ`, all measurable set `s`, + `∫⁻ a in s, ennreal.of_real (cond_cdf ρ a x) ∂ρ.fst = ρ (s ×ˢ Iic x)`. + +## References + +The construction of the conditional cdf in this file follows the proof of Theorem 3.4 in +[O. Kallenberg, Foundations of modern probability][kallenberg2021]. + +## TODO + +* The conditional cdf can be used to define the cdf of a real measure by using the + conditional cdf of `(measure.dirac unit.star).prod μ : measure (unit × ℝ)`. + +-/ + +open measure_theory set filter topological_space + +open_locale nnreal ennreal measure_theory topology probability_theory + +section aux_lemmas_to_be_moved + +variables {α β ι : Type*} + +namespace directed +-- todo after the port: move this to logic.encodable.basic near sequence_mono +variables [encodable α] [inhabited α] [preorder β] {f : α → β} (hf : directed (≥) f) + +lemma sequence_anti : antitone (f ∘ (hf.sequence f)) := +antitone_nat_of_succ_le $ hf.sequence_mono_nat + +lemma sequence_le (a : α) : f (hf.sequence f (encodable.encode a + 1)) ≤ f a := +hf.rel_sequence a + +end directed + +-- todo: move to data/set/lattice next to prod_Union or prod_sInter +lemma prod_Inter {s : set α} {t : ι → set β} [hι : nonempty ι] : + s ×ˢ (⋂ i, t i) = ⋂ i, s ×ˢ (t i) := +begin + ext x, + simp only [mem_prod, mem_Inter], + exact ⟨λ h i, ⟨h.1, h.2 i⟩, λ h, ⟨(h hι.some).1, λ i, (h i).2⟩⟩, +end + +lemma real.Union_Iic_rat : (⋃ r : ℚ, Iic (r : ℝ)) = univ := +begin + ext1, + simp only [mem_Union, mem_Iic, mem_univ, iff_true], + obtain ⟨r, hr⟩ := exists_rat_gt x, + exact ⟨r, hr.le⟩, +end + +lemma real.Inter_Iic_rat : (⋂ r : ℚ, Iic (r : ℝ)) = ∅ := +begin + ext1, + simp only [mem_Inter, mem_Iic, mem_empty_iff_false, iff_false, not_forall, not_le], + exact exists_rat_lt x, +end + +-- todo after the port: move to order/filter/at_top_bot +lemma at_bot_le_nhds_bot {α : Type*} [topological_space α] [linear_order α] [order_bot α] + [order_topology α] : + (at_bot : filter α) ≤ 𝓝 ⊥ := +begin + casesI subsingleton_or_nontrivial α, + { simp only [nhds_discrete, le_pure_iff, mem_at_bot_sets, mem_singleton_iff, + eq_iff_true_of_subsingleton, implies_true_iff, exists_const], }, + have h : at_bot.has_basis (λ _ : α, true) Iic := @at_bot_basis α _ _, + have h_nhds : (𝓝 ⊥).has_basis (λ a : α, ⊥ < a) (λ a, Iio a) := @nhds_bot_basis α _ _ _ _ _, + intro s, + rw [h.mem_iff, h_nhds.mem_iff], + rintros ⟨a, ha_bot_lt, h_Iio_a_subset_s⟩, + refine ⟨⊥, trivial, subset_trans _ h_Iio_a_subset_s⟩, + simpa only [Iic_bot, singleton_subset_iff, mem_Iio], +end + +-- todo after the port: move to order/filter/at_top_bot +lemma at_top_le_nhds_top {α : Type*} [topological_space α] [linear_order α] [order_top α] + [order_topology α] : + (at_top : filter α) ≤ 𝓝 ⊤ := +@at_bot_le_nhds_bot αᵒᵈ _ _ _ _ + +-- todo: move to topology/algebra/order/monotone_convergence +lemma tendsto_of_antitone {ι α : Type*} [preorder ι] [topological_space α] + [conditionally_complete_linear_order α] [order_topology α] {f : ι → α} (h_mono : antitone f) : + tendsto f at_top at_bot ∨ (∃ l, tendsto f at_top (𝓝 l)) := +@tendsto_of_monotone ι αᵒᵈ _ _ _ _ _ h_mono + +-- todo: move to data/real/ennreal +lemma ennreal.of_real_cinfi (f : α → ℝ) [nonempty α] : + ennreal.of_real (⨅ i, f i) = ⨅ i, ennreal.of_real (f i) := +begin + by_cases hf : bdd_below (range f), + { exact monotone.map_cinfi_of_continuous_at ennreal.continuous_of_real.continuous_at + (λ i j hij, ennreal.of_real_le_of_real hij) hf, }, + { symmetry, + rw [real.infi_of_not_bdd_below hf, ennreal.of_real_zero, ← ennreal.bot_eq_zero, infi_eq_bot], + obtain ⟨y, hy_mem, hy_neg⟩ := not_bdd_below_iff.mp hf 0, + obtain ⟨i, rfl⟩ := mem_range.mpr hy_mem, + refine λ x hx, ⟨i, _⟩, + rwa ennreal.of_real_of_nonpos hy_neg.le, }, +end + +-- todo: move to measure_theory/measurable_space +/-- Monotone convergence for an infimum over a directed family and indexed by a countable type -/ +theorem lintegral_infi_directed_of_measurable {mα : measurable_space α} [countable β] + {f : β → α → ℝ≥0∞} {μ : measure α} (hμ : μ ≠ 0) + (hf : ∀ b, measurable (f b)) (hf_int : ∀ b, ∫⁻ a, f b a ∂μ ≠ ∞) (h_directed : directed (≥) f) : + ∫⁻ a, ⨅ b, f b a ∂μ = ⨅ b, ∫⁻ a, f b a ∂μ := +begin + casesI nonempty_encodable β, + casesI is_empty_or_nonempty β, + { simp only [with_top.cinfi_empty, lintegral_const], + rw [ennreal.top_mul, if_neg], + simp only [measure.measure_univ_eq_zero, hμ, not_false_iff], }, + inhabit β, + have : ∀ a, (⨅ b, f b a) = (⨅ n, f (h_directed.sequence f n) a), + { refine λ a, le_antisymm (le_infi (λ n, infi_le _ _)) + (le_infi (λ b, infi_le_of_le (encodable.encode b + 1) _)), + exact (h_directed.sequence_le b a), }, + calc ∫⁻ a, ⨅ b, f b a ∂μ + = ∫⁻ a, ⨅ n, f (h_directed.sequence f n) a ∂μ : by simp only [this] + ... = ⨅ n, ∫⁻ a, f (h_directed.sequence f n) a ∂μ : + by { rw lintegral_infi (λ n, _) h_directed.sequence_anti, + { exact hf_int _, }, + { exact hf _, }, } + ... = ⨅ b, ∫⁻ a, f b a ∂μ : + begin + refine le_antisymm (le_infi (λ b, _)) (le_infi (λ n, _)), + { exact infi_le_of_le (encodable.encode b + 1) (lintegral_mono $ h_directed.sequence_le b) }, + { exact infi_le (λb, ∫⁻ a, f b a ∂μ) _ }, + end +end + +-- todo: move to measure_theory/pi_system +lemma is_pi_system_Iic [semilattice_inf α] : @is_pi_system α (range Iic) := +by { rintros s ⟨us, rfl⟩ t ⟨ut, rfl⟩ _, rw [Iic_inter_Iic], exact ⟨us ⊓ ut, rfl⟩, } + +-- todo: move to measure_theory/pi_system +lemma is_pi_system_Ici [semilattice_sup α] : @is_pi_system α (range Ici) := +by { rintros s ⟨us, rfl⟩ t ⟨ut, rfl⟩ _, rw [Ici_inter_Ici], exact ⟨us ⊔ ut, rfl⟩, } + + +end aux_lemmas_to_be_moved + +namespace measure_theory.measure + +variables {α β : Type*} {mα : measurable_space α} (ρ : measure (α × ℝ)) + +include mα + +/-- Measure on `α` such that for a measurable set `s`, `ρ.Iic_snd r s = ρ (s ×ˢ Iic r)`. -/ +noncomputable +def Iic_snd (r : ℝ) : measure α := (ρ.restrict (univ ×ˢ Iic r)).fst + +lemma Iic_snd_apply (r : ℝ) {s : set α} (hs : measurable_set s) : + ρ.Iic_snd r s = ρ (s ×ˢ Iic r) := +by rw [Iic_snd, fst_apply hs, + restrict_apply' (measurable_set.univ.prod (measurable_set_Iic : measurable_set (Iic r))), + ← prod_univ, prod_inter_prod, inter_univ, univ_inter] + +lemma Iic_snd_univ (r : ℝ) : ρ.Iic_snd r univ = ρ (univ ×ˢ Iic r) := +Iic_snd_apply ρ r measurable_set.univ + +lemma Iic_snd_mono {r r' : ℝ} (h_le : r ≤ r') : ρ.Iic_snd r ≤ ρ.Iic_snd r' := +begin + intros s hs, + simp_rw Iic_snd_apply ρ _ hs, + refine measure_mono (prod_subset_prod_iff.mpr (or.inl ⟨subset_rfl, Iic_subset_Iic.mpr _⟩)), + exact_mod_cast h_le, +end + +lemma Iic_snd_le_fst (r : ℝ) : ρ.Iic_snd r ≤ ρ.fst := +begin + intros s hs, + simp_rw [fst_apply hs, Iic_snd_apply ρ r hs], + exact measure_mono (prod_subset_preimage_fst _ _), +end + +lemma Iic_snd_ac_fst (r : ℝ) : ρ.Iic_snd r ≪ ρ.fst := +measure.absolutely_continuous_of_le (Iic_snd_le_fst ρ r) + +lemma is_finite_measure.Iic_snd {ρ : measure (α × ℝ)} [is_finite_measure ρ] (r : ℝ) : + is_finite_measure (ρ.Iic_snd r) := +is_finite_measure_of_le _ (Iic_snd_le_fst ρ _) + +lemma infi_Iic_snd_gt (t : ℚ) {s : set α} (hs : measurable_set s) [is_finite_measure ρ] : + (⨅ r : {r' : ℚ // t < r'}, ρ.Iic_snd r s) = ρ.Iic_snd t s := +begin + simp_rw [ρ.Iic_snd_apply _ hs], + rw ← measure_Inter_eq_infi, + { rw ← prod_Inter, + congr' with x : 1, + simp only [mem_Inter, mem_Iic, subtype.forall, subtype.coe_mk], + refine ⟨λ h, _, λ h a hta, h.trans _⟩, + { refine le_of_forall_lt_rat_imp_le (λ q htq, h q _), + exact_mod_cast htq, }, + { exact_mod_cast hta.le, }, }, + { exact λ _, hs.prod measurable_set_Iic, }, + { refine monotone.directed_ge (λ r r' hrr', prod_subset_prod_iff.mpr (or.inl ⟨subset_rfl, _⟩)), + refine Iic_subset_Iic.mpr _, + simp_rw coe_coe, + exact_mod_cast hrr', }, + { exact ⟨⟨t+1, lt_add_one _⟩, measure_ne_top ρ _⟩, }, +end + +lemma tendsto_Iic_snd_at_top {s : set α} (hs : measurable_set s) : + tendsto (λ r : ℚ, ρ.Iic_snd r s) at_top (𝓝 (ρ.fst s)) := +begin + simp_rw [ρ.Iic_snd_apply _ hs, fst_apply hs, ← prod_univ], + rw [← real.Union_Iic_rat, prod_Union], + refine tendsto_measure_Union (λ r q hr_le_q x, _), + simp only [mem_prod, mem_Iic, and_imp], + refine λ hxs hxr, ⟨hxs, hxr.trans _⟩, + exact_mod_cast hr_le_q, +end + +lemma tendsto_Iic_snd_at_bot [is_finite_measure ρ] {s : set α} (hs : measurable_set s) : + tendsto (λ r : ℚ, ρ.Iic_snd r s) at_bot (𝓝 0) := +begin + simp_rw [ρ.Iic_snd_apply _ hs], + have h_empty : ρ (s ×ˢ ∅) = 0, by simp only [prod_empty, measure_empty], + rw [← h_empty, ← real.Inter_Iic_rat, prod_Inter], + suffices h_neg : tendsto (λ r : ℚ, ρ (s ×ˢ Iic (↑-r))) at_top (𝓝 (ρ (⋂ r : ℚ, s ×ˢ Iic (↑-r)))), + { have h_inter_eq : (⋂ r : ℚ, s ×ˢ Iic (↑-r)) = (⋂ r : ℚ, s ×ˢ Iic (r : ℝ)), + { ext1 x, + simp only [rat.cast_eq_id, id.def, mem_Inter, mem_prod, mem_Iic], + refine ⟨λ h i, ⟨(h i).1, _⟩, λ h i, ⟨(h i).1, _⟩⟩; have h' := h (-i), + { rw neg_neg at h', exact h'.2, }, + { exact h'.2, }, }, + rw h_inter_eq at h_neg, + have h_fun_eq : (λ (r : ℚ), ρ (s ×ˢ Iic (r : ℝ))) = (λ r, ρ (s ×ˢ Iic ↑(- -r))), + { simp_rw neg_neg, }, + rw h_fun_eq, + exact h_neg.comp tendsto_neg_at_bot_at_top, }, + refine tendsto_measure_Inter (λ q, hs.prod measurable_set_Iic) _ ⟨0, measure_ne_top ρ _⟩, + refine λ q r hqr, prod_subset_prod_iff.mpr (or.inl ⟨subset_rfl, λ x hx, _⟩), + simp only [rat.cast_neg, mem_Iic] at hx ⊢, + refine hx.trans (neg_le_neg _), + exact_mod_cast hqr, +end + +end measure_theory.measure + +open measure_theory + +namespace probability_theory + +variables {α β ι : Type*} {mα : measurable_space α} + +include mα + +local attribute [instance] measure_theory.measure.is_finite_measure.Iic_snd + +/-! ### Auxiliary definitions + +We build towards the definition of `probability_theory.cond_cdf`. We first define +`probability_theory.pre_cdf`, a function defined on `α × ℚ` with the properties of a cdf almost +everywhere. We then introduce `probability_theory.cond_cdf_rat`, a function on `α × ℚ` which has +the properties of a cdf for all `a : α`. We finally extend to `ℝ`. -/ + +/-- `pre_cdf` is the Radon-Nikodym derivative of `ρ.Iic_snd` with respect to `ρ.fst` at each +`r : ℚ`. This function `ℚ → α → ℝ≥0∞` is such that for almost all `a : α`, the function `ℚ → ℝ≥0∞` +satisfies the properties of a cdf (monotone with limit 0 at -∞ and 1 at +∞, right-continuous). + +We define this function on `ℚ` and not `ℝ` because `ℚ` is countable, which allows us to prove +properties of the form `∀ᵐ a ∂ρ.fst, ∀ q, P (pre_cdf q a)`, instead of the weaker +`∀ q, ∀ᵐ a ∂ρ.fst, P (pre_cdf q a)`. -/ +noncomputable +def pre_cdf (ρ : measure (α × ℝ)) (r : ℚ) : α → ℝ≥0∞ := measure.rn_deriv (ρ.Iic_snd r) ρ.fst + +lemma measurable_pre_cdf {ρ : measure (α × ℝ)} {r : ℚ} : measurable (pre_cdf ρ r) := +measure.measurable_rn_deriv _ _ + +lemma with_density_pre_cdf (ρ : measure (α × ℝ)) (r : ℚ) [is_finite_measure ρ] : + ρ.fst.with_density (pre_cdf ρ r) = ρ.Iic_snd r := +measure.absolutely_continuous_iff_with_density_rn_deriv_eq.mp (measure.Iic_snd_ac_fst ρ r) + +lemma set_lintegral_pre_cdf_fst (ρ : measure (α × ℝ)) (r : ℚ) {s : set α} + (hs : measurable_set s) [is_finite_measure ρ] : + ∫⁻ x in s, pre_cdf ρ r x ∂ρ.fst = ρ.Iic_snd r s := +begin + have : ∀ r, ∫⁻ x in s, pre_cdf ρ r x ∂ρ.fst = ∫⁻ x in s, (pre_cdf ρ r * 1) x ∂ρ.fst, + { simp only [mul_one, eq_self_iff_true, forall_const], }, + rw [this, ← set_lintegral_with_density_eq_set_lintegral_mul _ measurable_pre_cdf _ hs], + { simp only [with_density_pre_cdf ρ r, pi.one_apply, lintegral_one, measure.restrict_apply, + measurable_set.univ, univ_inter], }, + { rw (_ : (1 : α → ℝ≥0∞) = (λ _, 1)), + exacts [measurable_const, rfl], }, +end + +lemma monotone_pre_cdf (ρ : measure (α × ℝ)) [is_finite_measure ρ] : + ∀ᵐ a ∂ρ.fst, monotone (λ r, pre_cdf ρ r a) := +begin + simp_rw [monotone, ae_all_iff], + refine λ r r' hrr', ae_le_of_forall_set_lintegral_le_of_sigma_finite + measurable_pre_cdf measurable_pre_cdf (λ s hs hs_fin, _), + rw [set_lintegral_pre_cdf_fst ρ r hs, set_lintegral_pre_cdf_fst ρ r' hs], + refine measure.Iic_snd_mono ρ _ s hs, + exact_mod_cast hrr', +end + +lemma set_lintegral_infi_gt_pre_cdf (ρ : measure (α × ℝ)) [is_finite_measure ρ] (t : ℚ) + {s : set α} (hs : measurable_set s) : + ∫⁻ x in s, ⨅ r : Ioi t, pre_cdf ρ r x ∂ρ.fst = ρ.Iic_snd t s := +begin + refine le_antisymm _ _, + { have h : ∀ q : Ioi t, ∫⁻ x in s, ⨅ r : Ioi t, pre_cdf ρ r x ∂ρ.fst ≤ ρ.Iic_snd q s, + { intros q, + rw [coe_coe, ← set_lintegral_pre_cdf_fst ρ _ hs], + refine set_lintegral_mono_ae _ measurable_pre_cdf _, + { exact measurable_infi (λ _, measurable_pre_cdf), }, + { filter_upwards [monotone_pre_cdf] with a ha_mono, + exact λ _, infi_le _ q, }, }, + calc ∫⁻ x in s, (⨅ (r : Ioi t), pre_cdf ρ r x) ∂ρ.fst + ≤ ⨅ q : Ioi t, ρ.Iic_snd q s : le_infi h + ... = ρ.Iic_snd t s : measure.infi_Iic_snd_gt ρ t hs, }, + { rw (set_lintegral_pre_cdf_fst ρ t hs).symm, + refine set_lintegral_mono_ae measurable_pre_cdf _ _, + { exact measurable_infi (λ _, measurable_pre_cdf), }, + { filter_upwards [monotone_pre_cdf] with a ha_mono, + exact λ _, le_infi (λ r, ha_mono (le_of_lt r.prop)), }, }, +end + +lemma pre_cdf_le_one (ρ : measure (α × ℝ)) [is_finite_measure ρ] : + ∀ᵐ a ∂ρ.fst, ∀ r, pre_cdf ρ r a ≤ 1 := +begin + rw ae_all_iff, + refine λ r, ae_le_of_forall_set_lintegral_le_of_sigma_finite measurable_pre_cdf + measurable_const (λ s hs hs_fin, _), + rw set_lintegral_pre_cdf_fst ρ r hs, + simp only [pi.one_apply, lintegral_one, measure.restrict_apply, measurable_set.univ, univ_inter], + exact measure.Iic_snd_le_fst ρ r s hs, +end + +lemma tendsto_lintegral_pre_cdf_at_top (ρ : measure (α × ℝ)) [is_finite_measure ρ] : + tendsto (λ r, ∫⁻ a, pre_cdf ρ r a ∂ρ.fst) at_top (𝓝 (ρ univ)) := +begin + convert ρ.tendsto_Iic_snd_at_top measurable_set.univ, + { ext1 r, + rw [← set_lintegral_univ, set_lintegral_pre_cdf_fst ρ _ measurable_set.univ], }, + { exact measure.fst_univ.symm, }, +end + +lemma tendsto_lintegral_pre_cdf_at_bot (ρ : measure (α × ℝ)) [is_finite_measure ρ] : + tendsto (λ r, ∫⁻ a, pre_cdf ρ r a ∂ρ.fst) at_bot (𝓝 0) := +begin + convert ρ.tendsto_Iic_snd_at_bot measurable_set.univ, + ext1 r, + rw [← set_lintegral_univ, set_lintegral_pre_cdf_fst ρ _ measurable_set.univ], +end + +lemma tendsto_pre_cdf_at_top_one (ρ : measure (α × ℝ)) [is_finite_measure ρ] : + ∀ᵐ a ∂ρ.fst, tendsto (λ r, pre_cdf ρ r a) at_top (𝓝 1) := +begin + -- We show first that `pre_cdf` has a limit almost everywhere. That limit has to be at most 1. + -- We then show that the integral of `pre_cdf` tends to the integral of 1, and that it also tends + -- to the integral of the limit. Since the limit is at most 1 and has same integral as 1, it is + -- equal to 1 a.e. + have h_mono := monotone_pre_cdf ρ, + have h_le_one := pre_cdf_le_one ρ, + -- `pre_cdf` has a limit a.e. + have h_exists : ∀ᵐ a ∂ρ.fst, ∃ l, tendsto (λ r, pre_cdf ρ r a) at_top (𝓝 l), + { filter_upwards [h_mono, h_le_one] with a ha_mono ha_le_one, + have h_tendsto : tendsto (λ r, pre_cdf ρ r a) at_top at_top + ∨ ∃ l, tendsto (λ r, pre_cdf ρ r a) at_top (𝓝 l) := tendsto_of_monotone ha_mono, + cases h_tendsto with h_absurd h_tendsto, + { rw monotone.tendsto_at_top_at_top_iff ha_mono at h_absurd, + obtain ⟨r, hr⟩ := h_absurd 2, + exact absurd (hr.trans (ha_le_one r)) ennreal.one_lt_two.not_le, }, + { exact h_tendsto, }, }, + classical, + -- let `F` be the pointwise limit of `pre_cdf` where it exists, and 0 elsewhere. + let F : α → ℝ≥0∞ := λ a, + if h : ∃ l, tendsto (λ r, pre_cdf ρ r a) at_top (𝓝 l) then h.some else 0, + have h_tendsto_ℚ : ∀ᵐ a ∂ρ.fst, tendsto (λ r, pre_cdf ρ r a) at_top (𝓝 (F a)), + { filter_upwards [h_exists] with a ha, + simp_rw [F, dif_pos ha], + exact ha.some_spec }, + have h_tendsto_ℕ : ∀ᵐ a ∂ρ.fst, tendsto (λ n : ℕ, pre_cdf ρ n a) at_top (𝓝 (F a)), + { filter_upwards [h_tendsto_ℚ] with a ha using ha.comp tendsto_coe_nat_at_top_at_top, }, + have hF_ae_meas : ae_measurable F ρ.fst, + { refine ae_measurable_of_tendsto_metrizable_ae _ (λ n, _) h_tendsto_ℚ, + exact measurable_pre_cdf.ae_measurable, }, + have hF_le_one : ∀ᵐ a ∂ρ.fst, F a ≤ 1, + { filter_upwards [h_tendsto_ℚ, h_le_one] with a ha ha_le using le_of_tendsto' ha ha_le, }, + -- it suffices to show that the limit `F` is 1 a.e. + suffices : ∀ᵐ a ∂ρ.fst, F a = 1, + { filter_upwards [h_tendsto_ℚ, this] with a ha_tendsto ha_eq, + rwa ha_eq at ha_tendsto, }, + -- since `F` is at most 1, proving that its integral is the same as the integral of 1 will tell + -- us that `F` is 1 a.e. + have h_lintegral_eq : ∫⁻ a, F a ∂ρ.fst = ∫⁻ a, 1 ∂ρ.fst, + { have h_lintegral : tendsto (λ r : ℕ, ∫⁻ a, pre_cdf ρ r a ∂ρ.fst) at_top + (𝓝 (∫⁻ a, F a ∂ρ.fst)), + { refine lintegral_tendsto_of_tendsto_of_monotone -- does this exist only for ℕ? + (λ _, measurable_pre_cdf.ae_measurable) _ h_tendsto_ℕ, + filter_upwards [h_mono] with a ha, + refine λ n m hnm, ha _, + exact_mod_cast hnm, }, + have h_lintegral' : tendsto (λ r : ℕ, ∫⁻ a, pre_cdf ρ r a ∂ρ.fst) at_top + (𝓝 (∫⁻ a, 1 ∂ρ.fst)), + { rw [lintegral_one, measure.fst_univ], + exact (tendsto_lintegral_pre_cdf_at_top ρ).comp tendsto_coe_nat_at_top_at_top, }, + exact tendsto_nhds_unique h_lintegral h_lintegral', }, + have : ∫⁻ a, (1 - F a) ∂ρ.fst = 0, + { rw [lintegral_sub' hF_ae_meas _ hF_le_one, h_lintegral_eq, tsub_self], + calc ∫⁻ a, F a ∂ρ.fst = ∫⁻ a, 1 ∂ρ.fst : h_lintegral_eq + ... = ρ.fst univ : lintegral_one + ... = ρ univ : measure.fst_univ + ... ≠ ∞ : measure_ne_top ρ _, }, + rw lintegral_eq_zero_iff' (ae_measurable_const.sub hF_ae_meas) at this, + filter_upwards [this, hF_le_one] with ha h_one_sub_eq_zero h_le_one, + rw [pi.zero_apply, tsub_eq_zero_iff_le] at h_one_sub_eq_zero, + exact le_antisymm h_le_one h_one_sub_eq_zero, +end + +lemma tendsto_pre_cdf_at_bot_zero (ρ : measure (α × ℝ)) [is_finite_measure ρ] : + ∀ᵐ a ∂ρ.fst, tendsto (λ r, pre_cdf ρ r a) at_bot (𝓝 0) := +begin + -- We show first that `pre_cdf` has a limit in ℝ≥0∞ almost everywhere. + -- We then show that the integral of `pre_cdf` tends to 0, and that it also tends + -- to the integral of the limit. Since the limit is has integral 0, it is equal to 0 a.e. + suffices : ∀ᵐ a ∂ρ.fst, tendsto (λ r, pre_cdf ρ (-r) a) at_top (𝓝 0), + { filter_upwards [this] with a ha, + have h_eq_neg : (λ (r : ℚ), pre_cdf ρ r a) = (λ (r : ℚ), pre_cdf ρ (- -r) a), + { simp_rw neg_neg, }, + rw h_eq_neg, + exact ha.comp tendsto_neg_at_bot_at_top, }, + have h_exists : ∀ᵐ a ∂ρ.fst, ∃ l, tendsto (λ r, pre_cdf ρ (-r) a) at_top (𝓝 l), + { filter_upwards [monotone_pre_cdf ρ] with a ha, + have h_anti : antitone (λ r, pre_cdf ρ (-r) a) := λ p q hpq, ha (neg_le_neg hpq), + have h_tendsto : tendsto (λ r, pre_cdf ρ (-r) a) at_top at_bot + ∨ ∃ l, tendsto (λ r, pre_cdf ρ (-r) a) at_top (𝓝 l) := tendsto_of_antitone h_anti, + cases h_tendsto with h_bot h_tendsto, + { exact ⟨0, tendsto.mono_right h_bot at_bot_le_nhds_bot⟩, }, + { exact h_tendsto, }, }, + classical, + let F : α → ℝ≥0∞ := λ a, + if h : ∃ l, tendsto (λ r, pre_cdf ρ (-r) a) at_top (𝓝 l) then h.some else 0, + have h_tendsto : ∀ᵐ a ∂ρ.fst, tendsto (λ r, pre_cdf ρ (-r) a) at_top (𝓝 (F a)), + { filter_upwards [h_exists] with a ha, + simp_rw [F, dif_pos ha], + exact ha.some_spec, }, + suffices h_lintegral_eq : ∫⁻ a, F a ∂ρ.fst = 0, + { have hF_ae_meas : ae_measurable F ρ.fst, + { refine ae_measurable_of_tendsto_metrizable_ae _ (λ n, _) h_tendsto, + exact measurable_pre_cdf.ae_measurable, }, + rw [lintegral_eq_zero_iff' hF_ae_meas] at h_lintegral_eq, + filter_upwards [h_tendsto, h_lintegral_eq] with a ha_tendsto ha_eq, + rwa ha_eq at ha_tendsto, }, + have h_lintegral : tendsto (λ r, ∫⁻ a, pre_cdf ρ (-r) a ∂ρ.fst) at_top (𝓝 (∫⁻ a, F a ∂ρ.fst)), + { refine tendsto_lintegral_filter_of_dominated_convergence (λ _, 1) + (eventually_of_forall (λ _, measurable_pre_cdf)) (eventually_of_forall (λ _, _)) + _ h_tendsto, + { filter_upwards [pre_cdf_le_one ρ] with a ha using ha _, }, + { rw lintegral_one, + exact measure_ne_top _ _, }, }, + have h_lintegral' : tendsto (λ r, ∫⁻ a, pre_cdf ρ (-r) a ∂ρ.fst) at_top (𝓝 0), + { have h_lintegral_eq : (λ r, ∫⁻ a, pre_cdf ρ (-r) a ∂ρ.fst) = λ r, ρ (univ ×ˢ Iic (-r)), + { ext1 n, + rw [← set_lintegral_univ, set_lintegral_pre_cdf_fst ρ _ measurable_set.univ, + measure.Iic_snd_univ], + norm_cast, }, + rw h_lintegral_eq, + have h_zero_eq_measure_Inter : (0 : ℝ≥0∞) = ρ (⋂ r : ℚ, univ ×ˢ Iic (-r)), + { suffices : (⋂ r : ℚ, Iic (-(r : ℝ))) = ∅, + { rwa [← prod_Inter, this, prod_empty, measure_empty], }, + ext1 x, + simp only [mem_Inter, mem_Iic, mem_empty_iff_false, iff_false, not_forall, not_le], + simp_rw neg_lt, + exact exists_rat_gt _, }, + rw h_zero_eq_measure_Inter, + refine tendsto_measure_Inter (λ n, measurable_set.univ.prod measurable_set_Iic) + (λ i j hij x, _) ⟨0, measure_ne_top ρ _⟩, + simp only [mem_prod, mem_univ, mem_Iic, true_and], + refine λ hxj, hxj.trans (neg_le_neg _), + exact_mod_cast hij, }, + exact tendsto_nhds_unique h_lintegral h_lintegral', +end + +lemma inf_gt_pre_cdf (ρ : measure (α × ℝ)) [is_finite_measure ρ] : + ∀ᵐ a ∂ρ.fst, ∀ t : ℚ, (⨅ r : Ioi t, pre_cdf ρ r a) = pre_cdf ρ t a := +begin + rw ae_all_iff, + refine λ t, ae_eq_of_forall_set_lintegral_eq_of_sigma_finite _ measurable_pre_cdf _, + { exact measurable_infi (λ i, measurable_pre_cdf), }, + intros s hs hs_fin, + rw [set_lintegral_infi_gt_pre_cdf ρ t hs, set_lintegral_pre_cdf_fst ρ t hs], +end + + +section has_cond_cdf + +/-- A product measure on `α × ℝ` is said to have a conditional cdf at `a : α` if `pre_cdf` is +monotone with limit 0 at -∞ and 1 at +∞, and is right continuous. +This property holds almost everywhere (see `has_cond_cdf_ae`). -/ +structure has_cond_cdf (ρ : measure (α × ℝ)) (a : α) : Prop := +(mono : monotone (λ r, pre_cdf ρ r a)) +(le_one : ∀ r, pre_cdf ρ r a ≤ 1) +(tendsto_at_top_one : tendsto (λ r, pre_cdf ρ r a) at_top (𝓝 1)) +(tendsto_at_bot_zero : tendsto (λ r, pre_cdf ρ r a) at_bot (𝓝 0)) +(infi_rat_gt_eq : ∀ t : ℚ, (⨅ r : Ioi t, pre_cdf ρ r a) = pre_cdf ρ t a) + +lemma has_cond_cdf_ae (ρ : measure (α × ℝ)) [is_finite_measure ρ] : + ∀ᵐ a ∂ρ.fst, has_cond_cdf ρ a := +begin + filter_upwards [monotone_pre_cdf ρ, pre_cdf_le_one ρ, tendsto_pre_cdf_at_top_one ρ, + tendsto_pre_cdf_at_bot_zero ρ, inf_gt_pre_cdf ρ] with a h1 h2 h3 h4 h5, + exact ⟨h1, h2, h3, h4, h5⟩, +end + +/-- A measurable set of elements of `α` such that `ρ` has a conditional cdf at all +`a ∈ cond_cdf_set`. -/ +def cond_cdf_set (ρ : measure (α × ℝ)) : set α := (to_measurable ρ.fst {b | ¬ has_cond_cdf ρ b})ᶜ + +lemma measurable_set_cond_cdf_set (ρ : measure (α × ℝ)) : measurable_set (cond_cdf_set ρ) := +(measurable_set_to_measurable _ _).compl + +lemma has_cond_cdf_of_mem_cond_cdf_set {ρ : measure (α × ℝ)} {a : α} (h : a ∈ cond_cdf_set ρ) : + has_cond_cdf ρ a := +begin + rw [cond_cdf_set, mem_compl_iff] at h, + have h_ss := subset_to_measurable ρ.fst {b | ¬ has_cond_cdf ρ b}, + by_contra ha, + exact h (h_ss ha), +end + +lemma mem_cond_cdf_set_ae (ρ : measure (α × ℝ)) [is_finite_measure ρ] : + ∀ᵐ a ∂ρ.fst, a ∈ cond_cdf_set ρ := +begin + simp_rw [ae_iff, cond_cdf_set, not_mem_compl_iff, set_of_mem_eq, measure_to_measurable], + exact has_cond_cdf_ae ρ, +end + +end has_cond_cdf + + +open_locale classical + +/-- Conditional cdf of the measure given the value on `α`, restricted to the rationals. +It is defined to be `pre_cdf` if `a ∈ cond_cdf_set`, and a default cdf-like function +otherwise. This is an auxiliary definition used to define `cond_cdf`. -/ +noncomputable +def cond_cdf_rat (ρ : measure (α × ℝ)) : α → ℚ → ℝ := +λ a, if a ∈ cond_cdf_set ρ then (λ r, (pre_cdf ρ r a).to_real) else (λ r, if r < 0 then 0 else 1) + +lemma cond_cdf_rat_of_not_mem (ρ : measure (α × ℝ)) (a : α) (h : a ∉ cond_cdf_set ρ) {r : ℚ} : + cond_cdf_rat ρ a r = if r < 0 then 0 else 1 := +by simp only [cond_cdf_rat, h, if_false] + +lemma cond_cdf_rat_of_mem (ρ : measure (α × ℝ)) (a : α) (h : a ∈ cond_cdf_set ρ) (r : ℚ) : + cond_cdf_rat ρ a r = (pre_cdf ρ r a).to_real := +by simp only [cond_cdf_rat, h, if_true] + +lemma monotone_cond_cdf_rat (ρ : measure (α × ℝ)) (a : α) : + monotone (cond_cdf_rat ρ a) := +begin + by_cases h : a ∈ cond_cdf_set ρ, + { simp only [cond_cdf_rat, h, if_true, forall_const, and_self], + intros r r' hrr', + have h' := has_cond_cdf_of_mem_cond_cdf_set h, + have h_ne_top : ∀ r, pre_cdf ρ r a ≠ ∞ := λ r, ((h'.le_one r).trans_lt ennreal.one_lt_top).ne, + rw ennreal.to_real_le_to_real (h_ne_top _) (h_ne_top _), + exact h'.1 hrr', }, + { simp only [cond_cdf_rat, h, if_false], + intros x y hxy, + dsimp only, + split_ifs, + exacts [le_rfl, zero_le_one, absurd (hxy.trans_lt h_2) h_1, le_rfl], }, +end + +lemma measurable_cond_cdf_rat (ρ : measure (α × ℝ)) (q : ℚ) : + measurable (λ a, cond_cdf_rat ρ a q) := +begin + simp_rw [cond_cdf_rat, ite_apply], + exact measurable.ite (measurable_set_cond_cdf_set ρ) measurable_pre_cdf.ennreal_to_real + measurable_const, +end + +lemma cond_cdf_rat_nonneg (ρ : measure (α × ℝ)) (a : α) (r : ℚ) : + 0 ≤ cond_cdf_rat ρ a r := +by { unfold cond_cdf_rat, split_ifs, exacts [ennreal.to_real_nonneg, le_rfl, zero_le_one], } + +lemma cond_cdf_rat_le_one (ρ : measure (α × ℝ)) (a : α) (r : ℚ) : + cond_cdf_rat ρ a r ≤ 1 := +begin + unfold cond_cdf_rat, + split_ifs with h, + { refine ennreal.to_real_le_of_le_of_real zero_le_one _, + rw ennreal.of_real_one, + exact (has_cond_cdf_of_mem_cond_cdf_set h).le_one r, }, + exacts [zero_le_one, le_rfl], +end + +lemma tendsto_cond_cdf_rat_at_bot (ρ : measure (α × ℝ)) (a : α) : + tendsto (cond_cdf_rat ρ a) at_bot (𝓝 0) := +begin + unfold cond_cdf_rat, + split_ifs with h, + { rw [← ennreal.zero_to_real, ennreal.tendsto_to_real_iff], + { exact (has_cond_cdf_of_mem_cond_cdf_set h).tendsto_at_bot_zero, }, + { have h' := has_cond_cdf_of_mem_cond_cdf_set h, + exact λ r, ((h'.le_one r).trans_lt ennreal.one_lt_top).ne, }, + { exact ennreal.zero_ne_top, }, }, + { refine (tendsto_congr' _).mp tendsto_const_nhds, + rw [eventually_eq, eventually_at_bot], + refine ⟨-1, λ q hq, (if_pos (hq.trans_lt _)).symm⟩, + linarith, }, +end + +lemma tendsto_cond_cdf_rat_at_top (ρ : measure (α × ℝ)) (a : α) : + tendsto (cond_cdf_rat ρ a) at_top (𝓝 1) := +begin + unfold cond_cdf_rat, + split_ifs with h, + { have h' := has_cond_cdf_of_mem_cond_cdf_set h, + rw [← ennreal.one_to_real, ennreal.tendsto_to_real_iff], + { exact h'.tendsto_at_top_one, }, + { exact λ r, ((h'.le_one r).trans_lt ennreal.one_lt_top).ne, }, + { exact ennreal.one_ne_top, }, }, + { refine (tendsto_congr' _).mp tendsto_const_nhds, + rw [eventually_eq, eventually_at_top], + exact ⟨0, λ q hq, (if_neg (not_lt.mpr hq)).symm⟩, }, +end + +lemma cond_cdf_rat_ae_eq (ρ : measure (α × ℝ)) [is_finite_measure ρ] (r : ℚ) : + (λ a, cond_cdf_rat ρ a r) =ᵐ[ρ.fst] λ a, (pre_cdf ρ r a).to_real := +by filter_upwards [mem_cond_cdf_set_ae ρ] with a ha using cond_cdf_rat_of_mem ρ a ha r + +lemma of_real_cond_cdf_rat_ae_eq (ρ : measure (α × ℝ)) [is_finite_measure ρ] (r : ℚ) : + (λ a, ennreal.of_real (cond_cdf_rat ρ a r)) =ᵐ[ρ.fst] pre_cdf ρ r := +begin + filter_upwards [cond_cdf_rat_ae_eq ρ r, pre_cdf_le_one ρ] with a ha ha_le_one, + rw [ha, ennreal.of_real_to_real], + exact ((ha_le_one r).trans_lt ennreal.one_lt_top).ne, +end + +lemma inf_gt_cond_cdf_rat (ρ : measure (α × ℝ)) (a : α) (t : ℚ) : + (⨅ r : Ioi t, cond_cdf_rat ρ a r) = cond_cdf_rat ρ a t := +begin + by_cases ha : a ∈ cond_cdf_set ρ, + { simp_rw cond_cdf_rat_of_mem ρ a ha, + have ha' := has_cond_cdf_of_mem_cond_cdf_set ha, + rw ← ennreal.to_real_infi, + { suffices : (⨅ (i : ↥(Ioi t)), pre_cdf ρ ↑i a) = pre_cdf ρ t a, by rw this, + rw ← ha'.infi_rat_gt_eq, }, + { exact λ r, ((ha'.le_one r).trans_lt ennreal.one_lt_top).ne, }, }, + { simp_rw cond_cdf_rat_of_not_mem ρ a ha, + have h_bdd : bdd_below (range (λ (r : ↥(Ioi t)), ite ((r : ℚ) < 0) (0 : ℝ) 1)), + { refine ⟨0, λ x hx, _⟩, + obtain ⟨y, rfl⟩ := mem_range.mpr hx, + dsimp only, + split_ifs, + exacts [le_rfl, zero_le_one], }, + split_ifs with h h, + { refine le_antisymm _ (le_cinfi (λ x, _)), + { obtain ⟨q, htq, hq_neg⟩ : ∃ q, t < q ∧ q < 0, + { refine ⟨t/2, _, _⟩, + { linarith, }, + { linarith, }, }, + refine (cinfi_le h_bdd ⟨q, htq⟩).trans _, + rw if_pos, + rwa subtype.coe_mk, }, + { split_ifs, + exacts [le_rfl, zero_le_one], }, }, + { refine le_antisymm _ _, + { refine (cinfi_le h_bdd ⟨t+1, lt_add_one t⟩).trans _, + split_ifs, + exacts [zero_le_one, le_rfl], }, + { refine le_cinfi (λ x, _), + rw if_neg, + rw not_lt at h ⊢, + exact h.trans (mem_Ioi.mp x.prop).le, }, }, }, +end + +/-- Conditional cdf of the measure given the value on `α`, as a plain function. This is an auxiliary +definition used to define `cond_cdf`. -/ +@[irreducible] noncomputable +def cond_cdf' (ρ : measure (α × ℝ)) : α → ℝ → ℝ := +λ a t, ⨅ r : {r' : ℚ // t < r'}, cond_cdf_rat ρ a r + +lemma cond_cdf'_def {ρ : measure (α × ℝ)} {a : α} {x : ℝ} : + cond_cdf' ρ a x = ⨅ r : {r : ℚ // x < r}, cond_cdf_rat ρ a r := +by rw cond_cdf' + +lemma cond_cdf'_eq_cond_cdf_rat (ρ : measure (α × ℝ)) (a : α) (r : ℚ) : + cond_cdf' ρ a r = cond_cdf_rat ρ a r := +begin + rw [← inf_gt_cond_cdf_rat ρ a r, cond_cdf'], + refine equiv.infi_congr _ _, + { exact + { to_fun := λ t, ⟨t.1, by exact_mod_cast t.2⟩, + inv_fun := λ t, ⟨t.1, by exact_mod_cast t.2⟩, + left_inv := λ t, by simp only [subtype.val_eq_coe, subtype.coe_eta], + right_inv := λ t, by simp only [subtype.val_eq_coe, subtype.coe_eta], }, }, + { intro t, + simp only [subtype.val_eq_coe, equiv.coe_fn_mk, subtype.coe_mk], }, +end + +lemma cond_cdf'_nonneg (ρ : measure (α × ℝ)) (a : α) (r : ℝ) : + 0 ≤ cond_cdf' ρ a r := +begin + haveI : nonempty {r' : ℚ // r < ↑r'}, + { obtain ⟨r, hrx⟩ := exists_rat_gt r, + exact ⟨⟨r, hrx⟩⟩, }, + rw cond_cdf'_def, + exact le_cinfi (λ r', cond_cdf_rat_nonneg ρ a _), +end + +lemma bdd_below_range_cond_cdf_rat_gt (ρ : measure (α × ℝ)) (a : α) (x : ℝ) : + bdd_below (range (λ (r : {r' : ℚ // x < ↑r'}), cond_cdf_rat ρ a r)) := +by { refine ⟨0, λ z, _⟩, rintros ⟨u, rfl⟩, exact cond_cdf_rat_nonneg ρ a _, } + +lemma monotone_cond_cdf' (ρ : measure (α × ℝ)) (a : α) : monotone (cond_cdf' ρ a) := +begin + intros x y hxy, + haveI : nonempty {r' : ℚ // y < ↑r'}, + { obtain ⟨r, hrx⟩ := exists_rat_gt y, + exact ⟨⟨r, hrx⟩⟩, }, + simp_rw cond_cdf'_def, + refine le_cinfi (λ r, (cinfi_le _ _).trans_eq _), + { exact ⟨r.1, hxy.trans_lt r.prop⟩, }, + { exact bdd_below_range_cond_cdf_rat_gt ρ a x, }, + { refl, }, +end + +lemma continuous_within_at_cond_cdf'_Ici (ρ : measure (α × ℝ)) (a : α) (x : ℝ) : + continuous_within_at (cond_cdf' ρ a) (Ici x) x := +begin + rw ← continuous_within_at_Ioi_iff_Ici, + convert monotone.tendsto_nhds_within_Ioi (monotone_cond_cdf' ρ a) x, + rw Inf_image', + have h' : (⨅ r : Ioi x, cond_cdf' ρ a r) = ⨅ r : {r' : ℚ // x < r'}, cond_cdf' ρ a r, + { refine infi_Ioi_eq_infi_rat_gt x _ (monotone_cond_cdf' ρ a), + refine ⟨0, λ z, _⟩, + rintros ⟨u, hux, rfl⟩, + exact cond_cdf'_nonneg ρ a u, }, + have h'' : (⨅ r : {r' : ℚ // x < r'}, cond_cdf' ρ a r) + = ⨅ r : {r' : ℚ // x < r'}, cond_cdf_rat ρ a r, + { congr' with r, + exact cond_cdf'_eq_cond_cdf_rat ρ a r, }, + rw [h', h'', continuous_within_at], + congr, + exact cond_cdf'_def, +end + +/-! ### Conditional cdf -/ + +/-- Conditional cdf of the measure given the value on `α`, as a Stieltjes function. -/ +noncomputable +def cond_cdf (ρ : measure (α × ℝ)) (a : α) : stieltjes_function := +{ to_fun := cond_cdf' ρ a, + mono' := monotone_cond_cdf' ρ a, + right_continuous' := λ x, continuous_within_at_cond_cdf'_Ici ρ a x, } + +lemma cond_cdf_eq_cond_cdf_rat (ρ : measure (α × ℝ)) (a : α) (r : ℚ) : + cond_cdf ρ a r = cond_cdf_rat ρ a r := +cond_cdf'_eq_cond_cdf_rat ρ a r + +/-- The conditional cdf is non-negative for all `a : α`. -/ +lemma cond_cdf_nonneg (ρ : measure (α × ℝ)) (a : α) (r : ℝ) : + 0 ≤ cond_cdf ρ a r := +cond_cdf'_nonneg ρ a r + +/-- The conditional cdf is lower or equal to 1 for all `a : α`. -/ +lemma cond_cdf_le_one (ρ : measure (α × ℝ)) (a : α) (x : ℝ) : + cond_cdf ρ a x ≤ 1 := +begin + obtain ⟨r, hrx⟩ := exists_rat_gt x, + rw ← stieltjes_function.infi_rat_gt_eq, + simp_rw [coe_coe, cond_cdf_eq_cond_cdf_rat], + refine cinfi_le_of_le (bdd_below_range_cond_cdf_rat_gt ρ a x) _ (cond_cdf_rat_le_one _ _ _), + exact ⟨r, hrx⟩, +end + +/-- The conditional cdf tends to 0 at -∞ for all `a : α`. -/ +lemma tendsto_cond_cdf_at_bot (ρ : measure (α × ℝ)) (a : α) : + tendsto (cond_cdf ρ a) at_bot (𝓝 0) := +begin + have h_exists : ∀ x : ℝ, ∃ q : ℚ, x < q ∧ ↑q < x + 1 := λ x, exists_rat_btwn (lt_add_one x), + let qs : ℝ → ℚ := λ x, (h_exists x).some, + have hqs_tendsto : tendsto qs at_bot at_bot, + { rw tendsto_at_bot_at_bot, + refine λ q, ⟨q - 1, λ y hy, _⟩, + have h_le : ↑(qs y) ≤ (q : ℝ) - 1 + 1 := + ((h_exists y).some_spec.2.le).trans (add_le_add hy le_rfl), + rw sub_add_cancel at h_le, + exact_mod_cast h_le, }, + refine tendsto_of_tendsto_of_tendsto_of_le_of_le tendsto_const_nhds + ((tendsto_cond_cdf_rat_at_bot ρ a).comp hqs_tendsto) (cond_cdf_nonneg ρ a) (λ x, _), + rw [function.comp_apply, ← cond_cdf_eq_cond_cdf_rat], + exact (cond_cdf ρ a).mono (h_exists x).some_spec.1.le, +end + +/-- The conditional cdf tends to 1 at +∞ for all `a : α`. -/ +lemma tendsto_cond_cdf_at_top (ρ : measure (α × ℝ)) (a : α) : + tendsto (cond_cdf ρ a) at_top (𝓝 1) := +begin + have h_exists : ∀ x : ℝ, ∃ q : ℚ, x-1 < q ∧ ↑q < x := λ x, exists_rat_btwn (sub_one_lt x), + let qs : ℝ → ℚ := λ x, (h_exists x).some, + have hqs_tendsto : tendsto qs at_top at_top, + { rw tendsto_at_top_at_top, + refine λ q, ⟨q + 1, λ y hy, _⟩, + have h_le : y - 1 ≤ qs y := (h_exists y).some_spec.1.le, + rw sub_le_iff_le_add at h_le, + exact_mod_cast le_of_add_le_add_right (hy.trans h_le),}, + refine tendsto_of_tendsto_of_tendsto_of_le_of_le + ((tendsto_cond_cdf_rat_at_top ρ a).comp hqs_tendsto) tendsto_const_nhds _ (cond_cdf_le_one ρ a), + intro x, + rw [function.comp_apply, ← cond_cdf_eq_cond_cdf_rat], + exact (cond_cdf ρ a).mono (le_of_lt (h_exists x).some_spec.2), +end + +lemma cond_cdf_ae_eq (ρ : measure (α × ℝ)) [is_finite_measure ρ] (r : ℚ) : + (λ a, cond_cdf ρ a r) =ᵐ[ρ.fst] λ a, (pre_cdf ρ r a).to_real := +by filter_upwards [mem_cond_cdf_set_ae ρ] with a ha + using (cond_cdf_eq_cond_cdf_rat ρ a r).trans (cond_cdf_rat_of_mem ρ a ha r) + +lemma of_real_cond_cdf_ae_eq (ρ : measure (α × ℝ)) [is_finite_measure ρ] (r : ℚ) : + (λ a, ennreal.of_real (cond_cdf ρ a r)) =ᵐ[ρ.fst] pre_cdf ρ r := +begin + filter_upwards [cond_cdf_ae_eq ρ r, pre_cdf_le_one ρ] with a ha ha_le_one, + rw [ha, ennreal.of_real_to_real], + exact ((ha_le_one r).trans_lt ennreal.one_lt_top).ne, +end + +/-- The conditional cdf is a measurable function of `a : α` for all `x : ℝ`. -/ +lemma measurable_cond_cdf (ρ : measure (α × ℝ)) (x : ℝ) : + measurable (λ a, cond_cdf ρ a x) := +begin + have : (λ a, cond_cdf ρ a x) = λ a, (⨅ (r : {r' // x < ↑r'}), cond_cdf_rat ρ a ↑r), + { ext1 a, + rw ← stieltjes_function.infi_rat_gt_eq, + congr' with q, + rw [coe_coe, cond_cdf_eq_cond_cdf_rat], }, + rw this, + exact measurable_cinfi (λ q, measurable_cond_cdf_rat ρ q) + (λ a, bdd_below_range_cond_cdf_rat_gt ρ a _), +end + +/-- Auxiliary lemma for `set_lintegral_cond_cdf`. -/ +lemma set_lintegral_cond_cdf_rat (ρ : measure (α × ℝ)) [is_finite_measure ρ] (r : ℚ) + {s : set α} (hs : measurable_set s) : + ∫⁻ a in s, ennreal.of_real (cond_cdf ρ a r) ∂ρ.fst = ρ (s ×ˢ Iic r) := +begin + have : ∀ᵐ a ∂ρ.fst, a ∈ s → ennreal.of_real (cond_cdf ρ a r) = pre_cdf ρ r a, + { filter_upwards [of_real_cond_cdf_ae_eq ρ r] with a ha using λ _, ha, }, + rw [set_lintegral_congr_fun hs this, set_lintegral_pre_cdf_fst ρ r hs], + exact ρ.Iic_snd_apply r hs, +end + +lemma set_lintegral_cond_cdf (ρ : measure (α × ℝ)) [is_finite_measure ρ] (x : ℝ) + {s : set α} (hs : measurable_set s) : + ∫⁻ a in s, ennreal.of_real (cond_cdf ρ a x) ∂ρ.fst = ρ (s ×ˢ Iic x) := +begin + -- We have the result for `x : ℚ` thanks to `set_lintegral_cond_cdf_rat`. We use the equality + -- `cond_cdf ρ a x = ⨅ r : {r' : ℚ // x < r'}, cond_cdf ρ a r` and a monotone convergence + -- argument to extend it to the reals. + by_cases hρ_zero : ρ.fst.restrict s = 0, + { rw [hρ_zero, lintegral_zero_measure], + refine le_antisymm (zero_le _) _, + calc ρ (s ×ˢ Iic x) + ≤ ρ (prod.fst ⁻¹' s) : measure_mono (prod_subset_preimage_fst s (Iic x)) + ... = ρ.fst s : by rw [measure.fst_apply hs] + ... = ρ.fst.restrict s univ : by rw measure.restrict_apply_univ + ... = 0 : by simp only [hρ_zero, measure.coe_zero, pi.zero_apply], }, + have h : ∫⁻ a in s, ennreal.of_real (cond_cdf ρ a x) ∂ρ.fst + = ∫⁻ a in s, ennreal.of_real (⨅ r : {r' : ℚ // x < r'}, cond_cdf ρ a r) ∂ρ.fst, + { congr' with a : 1, + rw ← (cond_cdf ρ a).infi_rat_gt_eq x, }, + haveI h_nonempty : nonempty {r' : ℚ // x < ↑r'}, + { obtain ⟨r, hrx⟩ := exists_rat_gt x, + exact ⟨⟨r, hrx⟩⟩, }, + rw h, + simp_rw ennreal.of_real_cinfi, + have h_coe : ∀ b : {r' : ℚ // x < ↑r'}, (b : ℝ) = ((b : ℚ) : ℝ) := λ _, by congr, + rw lintegral_infi_directed_of_measurable hρ_zero + (λ q : {r' : ℚ // x < ↑r'}, (measurable_cond_cdf ρ q).ennreal_of_real), + rotate, + { intro b, + simp_rw h_coe, + rw [set_lintegral_cond_cdf_rat ρ _ hs], + exact measure_ne_top ρ _, }, + { refine monotone.directed_ge (λ i j hij a, ennreal.of_real_le_of_real ((cond_cdf ρ a).mono _)), + rw [h_coe, h_coe], + exact_mod_cast hij, }, + simp_rw [h_coe, set_lintegral_cond_cdf_rat ρ _ hs], + rw ← measure_Inter_eq_infi, + { rw ← prod_Inter, + congr' with y, + simp only [mem_Inter, mem_Iic, subtype.forall, subtype.coe_mk], + exact ⟨le_of_forall_lt_rat_imp_le, λ hyx q hq, hyx.trans hq.le⟩, }, + { exact λ i, hs.prod measurable_set_Iic, }, + { refine monotone.directed_ge (λ i j hij, _), + refine prod_subset_prod_iff.mpr (or.inl ⟨subset_rfl, Iic_subset_Iic.mpr _⟩), + exact_mod_cast hij, }, + { exact ⟨h_nonempty.some, measure_ne_top _ _⟩, }, +end + +lemma lintegral_cond_cdf (ρ : measure (α × ℝ)) [is_finite_measure ρ] (x : ℝ) : + ∫⁻ a, ennreal.of_real (cond_cdf ρ a x) ∂ρ.fst = ρ (univ ×ˢ Iic x) := +by rw [← set_lintegral_univ, set_lintegral_cond_cdf ρ _ measurable_set.univ] + +/-- The conditional cdf is a strongly measurable function of `a : α` for all `x : ℝ`. -/ +lemma strongly_measurable_cond_cdf (ρ : measure (α × ℝ)) (x : ℝ) : + strongly_measurable (λ a, cond_cdf ρ a x) := +(measurable_cond_cdf ρ x).strongly_measurable + +lemma integrable_cond_cdf (ρ : measure (α × ℝ)) [is_finite_measure ρ] (x : ℝ) : + integrable (λ a, cond_cdf ρ a x) ρ.fst := +begin + refine integrable_of_forall_fin_meas_le _ (measure_lt_top ρ.fst univ) _ (λ t ht hρt, _), + { exact (strongly_measurable_cond_cdf ρ _).ae_strongly_measurable, }, + { have : ∀ y, (‖cond_cdf ρ y x‖₊ : ℝ≥0∞) ≤ 1, + { intro y, + rw real.nnnorm_of_nonneg (cond_cdf_nonneg _ _ _), + exact_mod_cast cond_cdf_le_one _ _ _, }, + refine (set_lintegral_mono (measurable_cond_cdf _ _).ennnorm + measurable_one (λ y _, this y)).trans _, + simp only [pi.one_apply, lintegral_one, measure.restrict_apply, measurable_set.univ, + univ_inter], + exact measure_mono (subset_univ _), }, +end + +lemma set_integral_cond_cdf (ρ : measure (α × ℝ)) [is_finite_measure ρ] (x : ℝ) + {s : set α} (hs : measurable_set s) : + ∫ a in s, cond_cdf ρ a x ∂ρ.fst = (ρ (s ×ˢ Iic x)).to_real := +begin + have h := set_lintegral_cond_cdf ρ x hs, + rw ← of_real_integral_eq_lintegral_of_real at h, + { rw [← h, ennreal.to_real_of_real], + exact integral_nonneg (λ _, cond_cdf_nonneg _ _ _), }, + { exact (integrable_cond_cdf _ _).integrable_on, }, + { exact eventually_of_forall (λ _, cond_cdf_nonneg _ _ _), }, +end + +lemma integral_cond_cdf (ρ : measure (α × ℝ)) [is_finite_measure ρ] (x : ℝ) : + ∫ a, cond_cdf ρ a x ∂ρ.fst = (ρ (univ ×ˢ Iic x)).to_real := +by rw [← set_integral_cond_cdf ρ _ measurable_set.univ, measure.restrict_univ] + +end probability_theory From 8ff51ea9f2f5875755582577883fc99db1cfab88 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 17 May 2023 12:41:52 +0000 Subject: [PATCH 1022/1291] feat(analysis/inner_product_space/pi_L2): norms of basis vectors (#19020) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This adds `‖euclidean_space.single i (a : 𝕜)‖ = ‖a‖` and other similar results. They hold more generally for `pi_Lp`, so they are proven there first. The statement of `linear_isometry_equiv.pi_Lp_congr_left_single` has also been corrected to include the missing typecast. --- src/analysis/inner_product_space/pi_L2.lean | 41 ++++++++++++--- src/analysis/normed_space/pi_Lp.lean | 57 ++++++++++++++++++++- 2 files changed, 90 insertions(+), 8 deletions(-) diff --git a/src/analysis/inner_product_space/pi_L2.lean b/src/analysis/inner_product_space/pi_L2.lean index 4c865e310c985..011794e4a6975 100644 --- a/src/analysis/inner_product_space/pi_L2.lean +++ b/src/analysis/inner_product_space/pi_L2.lean @@ -226,18 +226,45 @@ by simp [apply_ite conj] lemma euclidean_space.inner_single_right [decidable_eq ι] (i : ι) (a : 𝕜) (v : euclidean_space 𝕜 ι) : - ⟪v, euclidean_space.single i (a : 𝕜)⟫ = a * conj (v i) := + ⟪v, euclidean_space.single i (a : 𝕜)⟫ = a * conj (v i) := by simp [apply_ite conj, mul_comm] -lemma euclidean_space.pi_Lp_congr_left_single [decidable_eq ι] {ι' : Type*} [fintype ι'] - [decidable_eq ι'] (e : ι' ≃ ι) (i' : ι') : - linear_isometry_equiv.pi_Lp_congr_left 2 𝕜 𝕜 e (euclidean_space.single i' (1:𝕜)) = - euclidean_space.single (e i') (1:𝕜) := +@[simp] lemma euclidean_space.norm_single [decidable_eq ι] (i : ι) (a : 𝕜) : + ‖euclidean_space.single i (a : 𝕜)‖ = ‖a‖ := +(pi_Lp.norm_equiv_symm_single 2 (λ i, 𝕜) i a : _) + +@[simp] lemma euclidean_space.nnnorm_single [decidable_eq ι] (i : ι) (a : 𝕜) : + ‖euclidean_space.single i (a : 𝕜)‖₊ = ‖a‖₊ := +(pi_Lp.nnnorm_equiv_symm_single 2 (λ i, 𝕜) i a : _) + +@[simp] lemma euclidean_space.dist_single_same [decidable_eq ι] (i : ι) (a b : 𝕜) : + dist (euclidean_space.single i (a : 𝕜)) (euclidean_space.single i (b : 𝕜)) = dist a b := +(pi_Lp.dist_equiv_symm_single_same 2 (λ i, 𝕜) i a b : _) + +@[simp] lemma euclidean_space.nndist_single_same [decidable_eq ι] (i : ι) (a b : 𝕜) : + nndist (euclidean_space.single i (a : 𝕜)) (euclidean_space.single i (b : 𝕜)) = nndist a b := +(pi_Lp.nndist_equiv_symm_single_same 2 (λ i, 𝕜) i a b : _) + +@[simp] lemma euclidean_space.edist_single_same [decidable_eq ι] (i : ι) (a b : 𝕜) : + edist (euclidean_space.single i (a : 𝕜)) (euclidean_space.single i (b : 𝕜)) = edist a b := +(pi_Lp.edist_equiv_symm_single_same 2 (λ i, 𝕜) i a b : _) + +/-- `euclidean_space.single` forms an orthonormal family. -/ +lemma euclidean_space.orthonormal_single [decidable_eq ι] : + orthonormal 𝕜 (λ i : ι, euclidean_space.single i (1 : 𝕜)) := begin - ext i, - simpa using if_congr e.symm_apply_eq rfl rfl + simp_rw [orthonormal_iff_ite, euclidean_space.inner_single_left, map_one, one_mul, + euclidean_space.single_apply], + intros i j, + refl, end +lemma euclidean_space.pi_Lp_congr_left_single [decidable_eq ι] {ι' : Type*} [fintype ι'] + [decidable_eq ι'] (e : ι' ≃ ι) (i' : ι') (v : 𝕜): + linear_isometry_equiv.pi_Lp_congr_left 2 𝕜 𝕜 e (euclidean_space.single i' v) = + euclidean_space.single (e i') v := +linear_isometry_equiv.pi_Lp_congr_left_single e i' _ + variables (ι 𝕜 E) /-- An orthonormal basis on E is an identification of `E` with its dimensional-matching diff --git a/src/analysis/normed_space/pi_Lp.lean b/src/analysis/normed_space/pi_Lp.lean index d6ead6efaefb7..ec85d614f1aa0 100644 --- a/src/analysis/normed_space/pi_Lp.lean +++ b/src/analysis/normed_space/pi_Lp.lean @@ -623,7 +623,9 @@ linear_isometry_equiv.ext $ λ x, rfl @[simp] lemma _root_.linear_isometry_equiv.pi_Lp_congr_left_single [decidable_eq ι] [decidable_eq ι'] (e : ι ≃ ι') (i : ι) (v : E) : - linear_isometry_equiv.pi_Lp_congr_left p 𝕜 E e (pi.single i v) = pi.single (e i) v := + linear_isometry_equiv.pi_Lp_congr_left p 𝕜 E e ( + (pi_Lp.equiv p (λ _, E)).symm $ pi.single i v) = + (pi_Lp.equiv p (λ _, E)).symm (pi.single (e i) v) := begin funext x, simp [linear_isometry_equiv.pi_Lp_congr_left, linear_equiv.Pi_congr_left', equiv.Pi_congr_left', @@ -649,6 +651,59 @@ end @[simp] lemma equiv_symm_smul : (pi_Lp.equiv p β).symm (c • x') = c • (pi_Lp.equiv p β).symm x' := rfl +section single + +variables (p) +variables [decidable_eq ι] + +@[simp] +lemma nnnorm_equiv_symm_single (i : ι) (b : β i) : + ‖(pi_Lp.equiv p β).symm (pi.single i b)‖₊ = ‖b‖₊ := +begin + haveI : nonempty ι := ⟨i⟩, + unfreezingI { induction p using with_top.rec_top_coe }, + { simp_rw [nnnorm_eq_csupr, equiv_symm_apply], + refine csupr_eq_of_forall_le_of_forall_lt_exists_gt (λ j, _) (λ n hn, ⟨i, hn.trans_eq _⟩), + { obtain rfl | hij := decidable.eq_or_ne i j, + { rw pi.single_eq_same }, + { rw [pi.single_eq_of_ne' hij, nnnorm_zero], + exact zero_le _ } }, + { rw pi.single_eq_same } }, + { have hp0 : (p : ℝ) ≠ 0, + { exact_mod_cast (zero_lt_one.trans_le $ fact.out (1 ≤ (p : ℝ≥0∞))).ne' }, + rw [nnnorm_eq_sum ennreal.coe_ne_top, ennreal.coe_to_real, fintype.sum_eq_single i, + equiv_symm_apply, pi.single_eq_same, ←nnreal.rpow_mul, one_div, mul_inv_cancel hp0, + nnreal.rpow_one], + intros j hij, + rw [equiv_symm_apply, pi.single_eq_of_ne hij, nnnorm_zero, nnreal.zero_rpow hp0] }, +end + +@[simp] +lemma norm_equiv_symm_single (i : ι) (b : β i) : + ‖(pi_Lp.equiv p β).symm (pi.single i b)‖ = ‖b‖ := +congr_arg coe $ nnnorm_equiv_symm_single p β i b + +@[simp] +lemma nndist_equiv_symm_single_same (i : ι) (b₁ b₂ : β i) : + nndist ((pi_Lp.equiv p β).symm (pi.single i b₁)) ((pi_Lp.equiv p β).symm (pi.single i b₂)) = + nndist b₁ b₂ := +by rw [nndist_eq_nnnorm, nndist_eq_nnnorm, ←equiv_symm_sub, ←pi.single_sub, + nnnorm_equiv_symm_single] + +@[simp] +lemma dist_equiv_symm_single_same (i : ι) (b₁ b₂ : β i) : + dist ((pi_Lp.equiv p β).symm (pi.single i b₁)) ((pi_Lp.equiv p β).symm (pi.single i b₂)) = + dist b₁ b₂ := +congr_arg coe $ nndist_equiv_symm_single_same p β i b₁ b₂ + +@[simp] +lemma edist_equiv_symm_single_same (i : ι) (b₁ b₂ : β i) : + edist ((pi_Lp.equiv p β).symm (pi.single i b₁)) ((pi_Lp.equiv p β).symm (pi.single i b₂)) = + edist b₁ b₂ := +by simpa only [edist_nndist] using congr_arg coe (nndist_equiv_symm_single_same p β i b₁ b₂) + +end single + /-- When `p = ∞`, this lemma does not hold without the additional assumption `nonempty ι` because the left-hand side simplifies to `0`, while the right-hand side simplifies to `‖b‖₊`. See `pi_Lp.nnnorm_equiv_symm_const'` for a version which exchanges the hypothesis `p ≠ ∞` for From 17fe3632366bfefa54c240db521ce21beeb7a28a Mon Sep 17 00:00:00 2001 From: Floris van Doorn Date: Wed, 17 May 2023 15:12:57 +0000 Subject: [PATCH 1023/1291] feat(geometry/manifold/cont_mdiff_mfderiv): prove that mfderiv is smooth (#18827) * From the sphere eversion project * `cont_mdiff_at.mfderiv` is strong enough to prove `cont_mdiff.cont_mdiff_tangent_map`. This is already done in [the sphere eversion project](https://github.com/leanprover-community/sphere-eversion/blob/9486e7f42c283041bc512f133daec359b73d4986/src/to_mathlib/unused/geometry_manifold_misc.lean#L383). We would need to generalize `cont_mdiff_at.mfderiv` to something like `cont_mdiff_within_at.mfderiv_within` to prove the full version of `cont_mdiff_on.cont_mdiff_on_tangent_map_within`. This is non-trivial and needs additions to already ported files, so I'll wait with doing that till after the port. --- src/geometry/manifold/cont_mdiff_mfderiv.lean | 120 ++++++++++++++++++ .../manifold/vector_bundle/tangent.lean | 49 ++++++- 2 files changed, 168 insertions(+), 1 deletion(-) diff --git a/src/geometry/manifold/cont_mdiff_mfderiv.lean b/src/geometry/manifold/cont_mdiff_mfderiv.lean index b002d3ee823d9..3da2e58fd45f5 100644 --- a/src/geometry/manifold/cont_mdiff_mfderiv.lean +++ b/src/geometry/manifold/cont_mdiff_mfderiv.lean @@ -94,6 +94,126 @@ lemma smooth.mdifferentiable_within_at (hf : smooth I I' f) : mdifferentiable_within_at I I' f s x := hf.mdifferentiable_at.mdifferentiable_within_at +/-! ### The derivative of a smooth function is smooth -/ + +section mfderiv + +include Is I's Js + +/-- The function that sends `x` to the `y`-derivative of `f(x,y)` at `g(x)` is `C^m` at `x₀`, +where the derivative is taken as a continuous linear map. +We have to assume that `f` is `C^n` at `(x₀, g(x₀))` for `n ≥ m + 1` and `g` is `C^m` at `x₀`. +We have to insert a coordinate change from `x₀` to `x` to make the derivative sensible. +This result is used to show that maps into the 1-jet bundle and cotangent bundle are smooth. +`cont_mdiff_at.mfderiv_id` and `cont_mdiff_at.mfderiv_const` are special cases of this. + +This result should be generalized to a `cont_mdiff_within_at` for `mfderiv_within`. +If we do that, we can deduce `cont_mdiff_on.cont_mdiff_on_tangent_map_within` from this. +-/ +theorem cont_mdiff_at.mfderiv {x₀ : N} (f : N → M → M') (g : N → M) + (hf : cont_mdiff_at (J.prod I) I' n (function.uncurry f) (x₀, g x₀)) + (hg : cont_mdiff_at J I m g x₀) (hmn : m + 1 ≤ n) : + cont_mdiff_at J 𝓘(𝕜, E →L[𝕜] E') m + (in_tangent_coordinates I I' g (λ x, f x (g x)) (λ x, mfderiv I I' (f x) (g x)) x₀) x₀ := +begin + have h4f : continuous_at (λ x, f x (g x)) x₀, + { apply continuous_at.comp (by apply hf.continuous_at) (continuous_at_id.prod hg.continuous_at) }, + have h4f := h4f.preimage_mem_nhds (ext_chart_at_source_mem_nhds I' (f x₀ (g x₀))), + have h3f := cont_mdiff_at_iff_cont_mdiff_at_nhds.mp (hf.of_le $ (self_le_add_left 1 m).trans hmn), + have h2f : ∀ᶠ x₂ in 𝓝 x₀, cont_mdiff_at I I' 1 (f x₂) (g x₂), + { refine ((continuous_at_id.prod hg.continuous_at).tendsto.eventually h3f).mono (λ x hx, _), + exact hx.comp (g x) (cont_mdiff_at_const.prod_mk cont_mdiff_at_id) }, + have h2g := hg.continuous_at.preimage_mem_nhds (ext_chart_at_source_mem_nhds I (g x₀)), + have : cont_diff_within_at 𝕜 m (λ x, fderiv_within 𝕜 + (ext_chart_at I' (f x₀ (g x₀)) ∘ f ((ext_chart_at J x₀).symm x) ∘ (ext_chart_at I (g x₀)).symm) + (range I) (ext_chart_at I (g x₀) (g ((ext_chart_at J x₀).symm x)))) + (range J) (ext_chart_at J x₀ x₀), + { rw [cont_mdiff_at_iff] at hf hg, + simp_rw [function.comp, uncurry, ext_chart_at_prod, local_equiv.prod_coe_symm, + model_with_corners.range_prod] at hf ⊢, + refine cont_diff_within_at.fderiv_within _ hg.2 I.unique_diff hmn (mem_range_self _) _, + { simp_rw [ext_chart_at_to_inv], exact hf.2 }, + { rw [← image_subset_iff], + rintros _ ⟨x, hx, rfl⟩, + exact mem_range_self _ } }, + have : cont_mdiff_at J 𝓘(𝕜, E →L[𝕜] E') m + (λ x, fderiv_within 𝕜 (ext_chart_at I' (f x₀ (g x₀)) ∘ f x ∘ (ext_chart_at I (g x₀)).symm) + (range I) (ext_chart_at I (g x₀) (g x))) x₀, + { simp_rw [cont_mdiff_at_iff_source_of_mem_source (mem_chart_source G x₀), + cont_mdiff_within_at_iff_cont_diff_within_at, function.comp], + exact this }, + have : cont_mdiff_at J 𝓘(𝕜, E →L[𝕜] E') m + (λ x, fderiv_within 𝕜 (ext_chart_at I' (f x₀ (g x₀)) ∘ (ext_chart_at I' (f x (g x))).symm ∘ + written_in_ext_chart_at I I' (g x) (f x) ∘ ext_chart_at I (g x) ∘ + (ext_chart_at I (g x₀)).symm) (range I) (ext_chart_at I (g x₀) (g x))) x₀, + { refine this.congr_of_eventually_eq _, + filter_upwards [h2g, h2f], + intros x₂ hx₂ h2x₂, + have : ∀ x ∈ (ext_chart_at I (g x₀)).symm ⁻¹' (ext_chart_at I (g x₂)).source ∩ + (ext_chart_at I (g x₀)).symm ⁻¹' (f x₂ ⁻¹' (ext_chart_at I' (f x₂ (g x₂))).source), + (ext_chart_at I' (f x₀ (g x₀)) ∘ (ext_chart_at I' (f x₂ (g x₂))).symm ∘ + written_in_ext_chart_at I I' (g x₂) (f x₂) ∘ ext_chart_at I (g x₂) ∘ + (ext_chart_at I (g x₀)).symm) x = + ext_chart_at I' (f x₀ (g x₀)) (f x₂ ((ext_chart_at I (g x₀)).symm x)), + { rintro x ⟨hx, h2x⟩, + simp_rw [written_in_ext_chart_at, function.comp_apply], + rw [(ext_chart_at I (g x₂)).left_inv hx, (ext_chart_at I' (f x₂ (g x₂))).left_inv h2x] }, + refine filter.eventually_eq.fderiv_within_eq_nhds (I.unique_diff _ $ mem_range_self _) _, + refine eventually_of_mem (inter_mem _ _) this, + { exact ext_chart_at_preimage_mem_nhds' _ _ hx₂ (ext_chart_at_source_mem_nhds I (g x₂)) }, + refine ext_chart_at_preimage_mem_nhds' _ _ hx₂ _, + exact (h2x₂.continuous_at).preimage_mem_nhds (ext_chart_at_source_mem_nhds _ _) }, + /- The conclusion is equal to the following, when unfolding coord_change of + `tangent_bundle_core` -/ + have : cont_mdiff_at J 𝓘(𝕜, E →L[𝕜] E') m + (λ x, (fderiv_within 𝕜 (ext_chart_at I' (f x₀ (g x₀)) ∘ (ext_chart_at I' (f x (g x))).symm) + (range I') (ext_chart_at I' (f x (g x)) (f x (g x)))).comp + ((mfderiv I I' (f x) (g x)).comp (fderiv_within 𝕜 (ext_chart_at I (g x) ∘ + (ext_chart_at I (g x₀)).symm) (range I) (ext_chart_at I (g x₀) (g x))))) x₀, + { refine this.congr_of_eventually_eq _, + filter_upwards [h2g, h2f, h4f], + intros x₂ hx₂ h2x₂ h3x₂, + symmetry, + rw [(h2x₂.mdifferentiable_at le_rfl).mfderiv], + have hI := (cont_diff_within_at_ext_coord_change I (g x₂) (g x₀) $ + local_equiv.mem_symm_trans_source _ hx₂ $ mem_ext_chart_source I (g x₂)) + .differentiable_within_at le_top, + have hI' := (cont_diff_within_at_ext_coord_change I' (f x₀ (g x₀)) (f x₂ (g x₂)) $ + local_equiv.mem_symm_trans_source _ + (mem_ext_chart_source I' (f x₂ (g x₂))) h3x₂).differentiable_within_at le_top, + have h3f := (h2x₂.mdifferentiable_at le_rfl).2, + refine fderiv_within.comp₃ _ hI' h3f hI _ _ _ _ (I.unique_diff _ $ mem_range_self _), + { exact λ x _, mem_range_self _ }, + { exact λ x _, mem_range_self _ }, + { simp_rw [written_in_ext_chart_at, function.comp_apply, + (ext_chart_at I (g x₂)).left_inv (mem_ext_chart_source I (g x₂))] }, + { simp_rw [function.comp_apply, (ext_chart_at I (g x₀)).left_inv hx₂] } }, + refine this.congr_of_eventually_eq _, + filter_upwards [h2g, h4f] with x hx h2x, + rw [in_tangent_coordinates_eq], + { refl }, + { rwa [ext_chart_at_source] at hx }, + { rwa [ext_chart_at_source] at h2x }, +end + +omit Js + +/-- The derivative `D_yf(y)` is `C^m` at `x₀`, where the derivative is taken as a continuous +linear map. We have to assume that `f` is `C^n` at `x₀` for some `n ≥ m + 1`. +We have to insert a coordinate change from `x₀` to `x` to make the derivative sensible. +This is a special case of `cont_mdiff_at.mfderiv` where `f` does not contain any parameters and +`g = id`. +-/ +lemma cont_mdiff_at.mfderiv_const {x₀ : M} {f : M → M'} + (hf : cont_mdiff_at I I' n f x₀) (hmn : m + 1 ≤ n) : + cont_mdiff_at I 𝓘(𝕜, E →L[𝕜] E') m (in_tangent_coordinates I I' id f (mfderiv I I' f) x₀) x₀ := +begin + have : cont_mdiff_at (I.prod I) I' n (λ x : M × M, f x.2) (x₀, x₀) := + cont_mdiff_at.comp (x₀, x₀) hf cont_mdiff_at_snd, + exact this.mfderiv (λ x, f) id cont_mdiff_at_id hmn, +end + +end mfderiv /-! ### The tangent map of a smooth function is smooth -/ diff --git a/src/geometry/manifold/vector_bundle/tangent.lean b/src/geometry/manifold/vector_bundle/tangent.lean index a707c965cf24e..bd35638c65230 100644 --- a/src/geometry/manifold/vector_bundle/tangent.lean +++ b/src/geometry/manifold/vector_bundle/tangent.lean @@ -29,16 +29,20 @@ This defines a smooth vector bundle `tangent_bundle` with fibers `tangent_space` bundle. -/ -open bundle set smooth_manifold_with_corners local_homeomorph +open bundle set smooth_manifold_with_corners local_homeomorph continuous_linear_map open_locale manifold topology bundle noncomputable theory variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] {E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] +{E' : Type*} [normed_add_comm_group E'] [normed_space 𝕜 E'] {H : Type*} [topological_space H] {I : model_with_corners 𝕜 E H} +{H' : Type*} [topological_space H'] {I' : model_with_corners 𝕜 E' H'} {M : Type*} [topological_space M] [charted_space H M] [smooth_manifold_with_corners I M] +{M' : Type*} [topological_space M'] [charted_space H' M'] [smooth_manifold_with_corners I' M'] {F : Type*} [normed_add_comm_group F] [normed_space 𝕜 F] + variables (I) /-- Auxiliary lemma for tangent spaces: the derivative of a coordinate change between two charts is @@ -342,3 +346,46 @@ rfl ((tangent_bundle_model_space_homeomorph H I).symm : model_prod H E → tangent_bundle I H) = (equiv.sigma_equiv_prod H E).symm := rfl + +section in_tangent_coordinates + +variables (I I') {M M' H H'} {N : Type*} + +/-- The map `in_coordinates` for the tangent bundle is trivial on the model spaces -/ +lemma in_coordinates_tangent_bundle_core_model_space + (x₀ x : H) (y₀ y : H') (ϕ : E →L[𝕜] E') : + in_coordinates E (tangent_space I) E' (tangent_space I') x₀ x y₀ y ϕ = ϕ := +begin + refine (vector_bundle_core.in_coordinates_eq _ _ _ _ _).trans _, + { exact mem_univ x }, + { exact mem_univ y }, + simp_rw [tangent_bundle_core_index_at, tangent_bundle_core_coord_change_model_space, + continuous_linear_map.id_comp, continuous_linear_map.comp_id] +end + +/-- When `ϕ x` is a continuous linear map that changes vectors in charts around `f x` to vectors +in charts around `g x`, `in_tangent_coordinates I I' f g ϕ x₀ x` is a coordinate change of +this continuous linear map that makes sense from charts around `f x₀` to charts around `g x₀` +by composing it with appropriate coordinate changes. +Note that the type of `ϕ` is more accurately +`Π x : N, tangent_space I (f x) →L[𝕜] tangent_space I' (g x)`. +We are unfolding `tangent_space` in this type so that Lean recognizes that the type of `ϕ` doesn't +actually depend on `f` or `g`. + +This is the underlying function of the trivializations of the hom of (pullbacks of) tangent spaces. +-/ +def in_tangent_coordinates (f : N → M) (g : N → M') (ϕ : N → E →L[𝕜] E') : N → N → E →L[𝕜] E' := +λ x₀ x, in_coordinates E (tangent_space I) E' (tangent_space I') (f x₀) (f x) (g x₀) (g x) (ϕ x) + +lemma in_tangent_coordinates_model_space (f : N → H) (g : N → H') (ϕ : N → E →L[𝕜] E') (x₀ : N) : + in_tangent_coordinates I I' f g ϕ x₀ = ϕ := +by simp_rw [in_tangent_coordinates, in_coordinates_tangent_bundle_core_model_space] + +lemma in_tangent_coordinates_eq (f : N → M) (g : N → M') (ϕ : N → E →L[𝕜] E') {x₀ x : N} + (hx : f x ∈ (chart_at H (f x₀)).source) (hy : g x ∈ (chart_at H' (g x₀)).source) : + in_tangent_coordinates I I' f g ϕ x₀ x = + (tangent_bundle_core I' M').coord_change (achart H' (g x)) (achart H' (g x₀)) (g x) ∘L ϕ x ∘L + (tangent_bundle_core I M).coord_change (achart H (f x₀)) (achart H (f x)) (f x) := +(tangent_bundle_core I M).in_coordinates_eq (tangent_bundle_core I' M') (ϕ x) hx hy + +end in_tangent_coordinates From c89fe2d59ae06402c3f55f978016d1ada444f57e Mon Sep 17 00:00:00 2001 From: Floris van Doorn Date: Wed, 17 May 2023 16:46:30 +0000 Subject: [PATCH 1024/1291] feat(geometry/manifold/vector_bundle/hom): the hom bundle is smooth (#18828) * From the sphere eversion project --- .../manifold/vector_bundle/basic.lean | 84 +++++++++++- src/geometry/manifold/vector_bundle/hom.lean | 122 ++++++++++++++++++ 2 files changed, 199 insertions(+), 7 deletions(-) create mode 100644 src/geometry/manifold/vector_bundle/hom.lean diff --git a/src/geometry/manifold/vector_bundle/basic.lean b/src/geometry/manifold/vector_bundle/basic.lean index 19ea2d529b202..0b38987317863 100644 --- a/src/geometry/manifold/vector_bundle/basic.lean +++ b/src/geometry/manifold/vector_bundle/basic.lean @@ -44,9 +44,13 @@ fields, they can also be C^k vector bundles, etc. * `bundle.total_space.smooth_manifold_with_corners`: A smooth vector bundle is naturally a smooth manifold. -* `vector_bundle_core.smooth_vector_bundle`: If a (topological) `vector_bundle_core` is smooth, in - the sense of having smooth transition functions, then the vector bundle constructed from it is a - smooth vector bundle. +* `vector_bundle_core.smooth_vector_bundle`: If a (topological) `vector_bundle_core` is smooth, + in the sense of having smooth transition functions (cf. `vector_bundle_core.is_smooth`), + then the vector bundle constructed from it is a smooth vector bundle. + +* `vector_prebundle.smooth_vector_bundle`: If a `vector_prebundle` is smooth, + in the sense of having smooth transition functions (cf. `vector_prebundle.is_smooth`), + then the vector bundle constructed from it is a smooth vector bundle. * `bundle.prod.smooth_vector_bundle`: The direct sum of two smooth vector bundles is a smooth vector bundle. @@ -237,10 +241,7 @@ end /-! ### Smooth vector bundles -/ -variables [nontrivially_normed_field 𝕜] [∀ x, add_comm_monoid (E x)] [∀ x, module 𝕜 (E x)] - [normed_add_comm_group F] [normed_space 𝕜 F] - [topological_space (total_space E)] [∀ x, topological_space (E x)] - +variables [nontrivially_normed_field 𝕜] {EB : Type*} [normed_add_comm_group EB] [normed_space 𝕜 EB] {HB : Type*} [topological_space HB] (IB : model_with_corners 𝕜 EB HB) [topological_space B] [charted_space HB B] [smooth_manifold_with_corners IB B] @@ -248,6 +249,12 @@ variables [nontrivially_normed_field 𝕜] [∀ x, add_comm_monoid (E x)] [∀ x {HM : Type*} [topological_space HM] {IM : model_with_corners 𝕜 EM HM} [topological_space M] [charted_space HM M] [Is : smooth_manifold_with_corners IM M] {n : ℕ∞} + [∀ x, add_comm_monoid (E x)] [∀ x, module 𝕜 (E x)] + [normed_add_comm_group F] [normed_space 𝕜 F] + +section with_topology + +variables [topological_space (total_space E)] [∀ x, topological_space (E x)] variables (F E) [fiber_bundle F E] [vector_bundle 𝕜 F E] @@ -380,3 +387,66 @@ instance bundle.prod.smooth_vector_bundle : end } end prod + +end with_topology + +/-! ### Prebundle construction for smooth vector bundles -/ + +namespace vector_prebundle + +variables {F E} + +/-- Mixin for a `vector_prebundle` stating smoothness of coordinate changes. -/ +class is_smooth (a : vector_prebundle 𝕜 F E) : Prop := +(exists_smooth_coord_change : ∀ (e e' ∈ a.pretrivialization_atlas), ∃ f : B → F →L[𝕜] F, + smooth_on IB 𝓘(𝕜, F →L[𝕜] F) f (e.base_set ∩ e'.base_set) ∧ + ∀ (b : B) (hb : b ∈ e.base_set ∩ e'.base_set) (v : F), + f b v = (e' (total_space_mk b (e.symm b v))).2) + +variables (a : vector_prebundle 𝕜 F E) [ha : a.is_smooth IB] {e e' : pretrivialization F (π E)} +include ha + +/-- A randomly chosen coordinate change on a `smooth_vector_prebundle`, given by + the field `exists_coord_change`. Note that `a.smooth_coord_change` need not be the same as + `a.coord_change`. -/ +noncomputable def smooth_coord_change (he : e ∈ a.pretrivialization_atlas) + (he' : e' ∈ a.pretrivialization_atlas) (b : B) : F →L[𝕜] F := +classical.some (ha.exists_smooth_coord_change e he e' he') b + +variables {IB} +lemma smooth_on_smooth_coord_change (he : e ∈ a.pretrivialization_atlas) + (he' : e' ∈ a.pretrivialization_atlas) : + smooth_on IB 𝓘(𝕜, F →L[𝕜] F) (a.smooth_coord_change IB he he') (e.base_set ∩ e'.base_set) := +(classical.some_spec (ha.exists_smooth_coord_change e he e' he')).1 + +lemma smooth_coord_change_apply (he : e ∈ a.pretrivialization_atlas) + (he' : e' ∈ a.pretrivialization_atlas) {b : B} (hb : b ∈ e.base_set ∩ e'.base_set) (v : F) : + a.smooth_coord_change IB he he' b v = (e' (total_space_mk b (e.symm b v))).2 := +(classical.some_spec (ha.exists_smooth_coord_change e he e' he')).2 b hb v + +lemma mk_smooth_coord_change (he : e ∈ a.pretrivialization_atlas) + (he' : e' ∈ a.pretrivialization_atlas) {b : B} (hb : b ∈ e.base_set ∩ e'.base_set) (v : F) : + (b, (a.smooth_coord_change IB he he' b v)) = e' (total_space_mk b (e.symm b v)) := +begin + ext, + { rw [e.mk_symm hb.1 v, e'.coe_fst', e.proj_symm_apply' hb.1], + rw [e.proj_symm_apply' hb.1], exact hb.2 }, + { exact a.smooth_coord_change_apply he he' hb v } +end + +variables (IB) +/-- Make a `smooth_vector_bundle` from a `smooth_vector_prebundle`. -/ +lemma smooth_vector_bundle : + @smooth_vector_bundle _ _ F E _ _ _ _ _ _ IB _ _ _ _ _ _ _ + a.total_space_topology a.fiber_topology a.to_fiber_bundle a.to_vector_bundle := +{ smooth_on_coord_change := begin + rintros _ _ ⟨e, he, rfl⟩ ⟨e', he', rfl⟩, + refine (a.smooth_on_smooth_coord_change he he').congr _, + intros b hb, + ext v, + rw [a.smooth_coord_change_apply he he' hb v, continuous_linear_equiv.coe_coe, + trivialization.coord_changeL_apply], + exacts [rfl, hb] + end } + +end vector_prebundle diff --git a/src/geometry/manifold/vector_bundle/hom.lean b/src/geometry/manifold/vector_bundle/hom.lean new file mode 100644 index 0000000000000..d81bcf54cbe50 --- /dev/null +++ b/src/geometry/manifold/vector_bundle/hom.lean @@ -0,0 +1,122 @@ +/- +Copyright (c) 2022 Floris van Doorn. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Floris van Doorn +-/ +import geometry.manifold.vector_bundle.basic +import topology.vector_bundle.hom + +/-! # Homs of smooth vector bundles over the same base space + +Here we show that `bundle.continuous_linear_map` is a smooth vector bundle. + +Note that we only do this for bundles of linear maps, not for bundles of arbitrary semilinear maps. +To do it for semilinear maps, we would need to generalize `continuous_linear_map.cont_mdiff` +(and `continuous_linear_map.cont_diff`) to semilinear maps. +-/ + +noncomputable theory + +open bundle set local_homeomorph continuous_linear_map pretrivialization +open_locale manifold bundle + +variables {𝕜 B F F₁ F₂ M M₁ M₂ : Type*} + {E : B → Type*} {E₁ : B → Type*} {E₂ : B → Type*} + [nontrivially_normed_field 𝕜] + [∀ x, add_comm_monoid (E x)] [∀ x, module 𝕜 (E x)] + [normed_add_comm_group F] [normed_space 𝕜 F] + [topological_space (total_space E)] [∀ x, topological_space (E x)] + [∀ x, add_comm_monoid (E₁ x)] [∀ x, module 𝕜 (E₁ x)] + [normed_add_comm_group F₁] [normed_space 𝕜 F₁] + [topological_space (total_space E₁)] [∀ x, topological_space (E₁ x)] + [∀ x, add_comm_monoid (E₂ x)] [∀ x, module 𝕜 (E₂ x)] + [normed_add_comm_group F₂] [normed_space 𝕜 F₂] + [topological_space (total_space E₂)] [∀ x, topological_space (E₂ x)] + + {EB : Type*} [normed_add_comm_group EB] [normed_space 𝕜 EB] + {HB : Type*} [topological_space HB] (IB : model_with_corners 𝕜 EB HB) + [topological_space B] [charted_space HB B] + {EM : Type*} [normed_add_comm_group EM] [normed_space 𝕜 EM] + {HM : Type*} [topological_space HM] {IM : model_with_corners 𝕜 EM HM} + [topological_space M] [charted_space HM M] [Is : smooth_manifold_with_corners IM M] + {n : ℕ∞} + [fiber_bundle F₁ E₁] [vector_bundle 𝕜 F₁ E₁] + [fiber_bundle F₂ E₂] [vector_bundle 𝕜 F₂ E₂] + {e₁ e₁' : trivialization F₁ (π E₁)} {e₂ e₂' : trivialization F₂ (π E₂)} + +local notation `LE₁E₂` := total_space (bundle.continuous_linear_map (ring_hom.id 𝕜) F₁ E₁ F₂ E₂) + +/- This proof is slow, especially the `simp only` and the elaboration of `h₂`. -/ +lemma smooth_on_continuous_linear_map_coord_change + [smooth_manifold_with_corners IB B] + [smooth_vector_bundle F₁ E₁ IB] [smooth_vector_bundle F₂ E₂ IB] + [mem_trivialization_atlas e₁] [mem_trivialization_atlas e₁'] + [mem_trivialization_atlas e₂] [mem_trivialization_atlas e₂'] : + smooth_on IB 𝓘(𝕜, ((F₁ →L[𝕜] F₂) →L[𝕜] (F₁ →L[𝕜] F₂))) + (continuous_linear_map_coord_change (ring_hom.id 𝕜) e₁ e₁' e₂ e₂') + ((e₁.base_set ∩ e₂.base_set) ∩ (e₁'.base_set ∩ e₂'.base_set)) := +begin + let L₁ := compL 𝕜 F₁ F₂ F₂, + have h₁ : smooth _ _ _ := L₁.cont_mdiff, + have h₂ : smooth _ _ _ := (continuous_linear_map.flip (compL 𝕜 F₁ F₁ F₂)).cont_mdiff, + have h₃ : smooth_on IB _ _ _ := smooth_on_coord_change e₁' e₁, + have h₄ : smooth_on IB _ _ _ := smooth_on_coord_change e₂ e₂', + refine ((h₁.comp_smooth_on (h₄.mono _)).clm_comp (h₂.comp_smooth_on (h₃.mono _))).congr _, + { mfld_set_tac }, + { mfld_set_tac }, + { intros b hb, ext L v, + simp only [continuous_linear_map_coord_change, continuous_linear_equiv.coe_coe, + continuous_linear_equiv.arrow_congrSL_apply, comp_apply, function.comp, compL_apply, + flip_apply, continuous_linear_equiv.symm_symm] }, +end + +variables [∀ x, has_continuous_add (E₂ x)] [∀ x, has_continuous_smul 𝕜 (E₂ x)] + +lemma hom_chart (y₀ y : LE₁E₂) : + chart_at (model_prod HB (F₁ →L[𝕜] F₂)) y₀ y = + (chart_at HB y₀.1 y.1, in_coordinates F₁ E₁ F₂ E₂ y₀.1 y.1 y₀.1 y.1 y.2) := +by simp_rw [fiber_bundle.charted_space_chart_at, trans_apply, local_homeomorph.prod_apply, + trivialization.coe_coe, local_homeomorph.refl_apply, function.id_def, hom_trivialization_at_apply] + +variables {IB} + +lemma cont_mdiff_at_hom_bundle (f : M → LE₁E₂) {x₀ : M} {n : ℕ∞} : + cont_mdiff_at IM (IB.prod 𝓘(𝕜, F₁ →L[𝕜] F₂)) n f x₀ ↔ + cont_mdiff_at IM IB n (λ x, (f x).1) x₀ ∧ + cont_mdiff_at IM 𝓘(𝕜, F₁ →L[𝕜] F₂) n + (λ x, in_coordinates F₁ E₁ F₂ E₂ (f x₀).1 (f x).1 (f x₀).1 (f x).1 (f x).2) x₀ := +by apply cont_mdiff_at_total_space + +lemma smooth_at_hom_bundle (f : M → LE₁E₂) {x₀ : M} : + smooth_at IM (IB.prod 𝓘(𝕜, F₁ →L[𝕜] F₂)) f x₀ ↔ + smooth_at IM IB (λ x, (f x).1) x₀ ∧ + smooth_at IM 𝓘(𝕜, F₁ →L[𝕜] F₂) + (λ x, in_coordinates F₁ E₁ F₂ E₂ (f x₀).1 (f x).1 (f x₀).1 (f x).1 (f x).2) x₀ := +cont_mdiff_at_hom_bundle f + +variables [smooth_manifold_with_corners IB B] + [smooth_vector_bundle F₁ E₁ IB] [smooth_vector_bundle F₂ E₂ IB] + +instance bundle.continuous_linear_map.vector_prebundle.is_smooth : + (bundle.continuous_linear_map.vector_prebundle (ring_hom.id 𝕜) F₁ E₁ F₂ E₂).is_smooth IB := +{ exists_smooth_coord_change := begin + rintro _ ⟨e₁, e₂, he₁, he₂, rfl⟩ _ ⟨e₁', e₂', he₁', he₂', rfl⟩, + resetI, + refine ⟨continuous_linear_map_coord_change (ring_hom.id 𝕜) e₁ e₁' e₂ e₂', + smooth_on_continuous_linear_map_coord_change IB, + continuous_linear_map_coord_change_apply (ring_hom.id 𝕜) e₁ e₁' e₂ e₂'⟩ + end } + +/-- Todo: remove this definition. It is probably needed because of the type-class pi bug +https://leanprover.zulipchat.com/#narrow/stream/116395-maths/topic/vector.20bundles.20--.20typeclass.20inference.20issue +-/ +@[reducible] +def smooth_vector_bundle.continuous_linear_map.aux (x) : + topological_space (bundle.continuous_linear_map (ring_hom.id 𝕜) F₁ E₁ F₂ E₂ x) := +by apply_instance +local attribute [instance, priority 1] smooth_vector_bundle.continuous_linear_map.aux + +instance smooth_vector_bundle.continuous_linear_map : + smooth_vector_bundle (F₁ →L[𝕜] F₂) (bundle.continuous_linear_map (ring_hom.id 𝕜) F₁ E₁ F₂ E₂) + IB := +(bundle.continuous_linear_map.vector_prebundle (ring_hom.id 𝕜) F₁ E₁ F₂ E₂).smooth_vector_bundle IB From a8b2226cfb0a79f5986492053fc49b1a0c6aeffb Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Wed, 17 May 2023 22:36:20 +0000 Subject: [PATCH 1025/1291] feat(analysis/convex/specific_functions): elementary convexity proofs (#19026) Give elementary proofs for the convexity of `pow`, `zpow`, `exp`, `log` and `rpow`, avoiding the second derivative test. See [Zulip](https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/Shortcut.20to.20integration) --- src/analysis/convex/slope.lean | 122 +++++++ src/analysis/convex/specific_functions.lean | 310 +++++++++++++----- src/analysis/special_functions/log/basic.lean | 8 + src/data/complex/exponential.lean | 122 ++++--- 4 files changed, 419 insertions(+), 143 deletions(-) diff --git a/src/analysis/convex/slope.lean b/src/analysis/convex/slope.lean index e6ddba269e7dc..b8e68c71b0c0e 100644 --- a/src/analysis/convex/slope.lean +++ b/src/analysis/convex/slope.lean @@ -217,6 +217,128 @@ lemma strict_concave_on_iff_slope_strict_anti_adjacent : ⟨λ h, ⟨h.1, λ x y z, h.slope_anti_adjacent⟩, λ h, strict_concave_on_of_slope_strict_anti_adjacent h.1 h.2⟩ +lemma convex_on.secant_mono_aux1 (hf : convex_on 𝕜 s f) + {x y z : 𝕜} (hx : x ∈ s) (hz : z ∈ s) (hxy : x < y) (hyz : y < z) : + (z - x) * f y ≤ (z - y) * f x + (y - x) * f z := +begin + have hxy' : 0 < y - x := by linarith, + have hyz' : 0 < z - y := by linarith, + have hxz' : 0 < z - x := by linarith, + rw ← le_div_iff' hxz', + have ha : 0 ≤ (z - y) / (z - x) := by positivity, + have hb : 0 ≤ (y - x) / (z - x) := by positivity, + calc f y = f ((z - y) / (z - x) * x + (y - x) / (z - x) * z) : _ + ... ≤ (z - y) / (z - x) * f x + (y - x) / (z - x) * f z : hf.2 hx hz ha hb _ + ... = ((z - y) * f x + (y - x) * f z) / (z - x) : _, + { congr' 1, + field_simp [hxy'.ne', hyz'.ne', hxz'.ne'], + ring }, + { field_simp [hxy'.ne', hyz'.ne', hxz'.ne'] }, + { field_simp [hxy'.ne', hyz'.ne', hxz'.ne'] } +end + +lemma convex_on.secant_mono_aux2 (hf : convex_on 𝕜 s f) + {x y z : 𝕜} (hx : x ∈ s) (hz : z ∈ s) (hxy : x < y) (hyz : y < z) : + (f y - f x) / (y - x) ≤ (f z - f x) / (z - x) := +begin + have hxy' : 0 < y - x := by linarith, + have hxz' : 0 < z - x := by linarith, + rw div_le_div_iff hxy' hxz', + linarith only [hf.secant_mono_aux1 hx hz hxy hyz], +end + +lemma convex_on.secant_mono_aux3 (hf : convex_on 𝕜 s f) + {x y z : 𝕜} (hx : x ∈ s) (hz : z ∈ s) (hxy : x < y) (hyz : y < z) : + (f z - f x) / (z - x) ≤ (f z - f y) / (z - y) := +begin + have hyz' : 0 < z - y := by linarith, + have hxz' : 0 < z - x := by linarith, + rw div_le_div_iff hxz' hyz', + linarith only [hf.secant_mono_aux1 hx hz hxy hyz], +end + +lemma convex_on.secant_mono (hf : convex_on 𝕜 s f) + {a x y : 𝕜} (ha : a ∈ s) (hx : x ∈ s) (hy : y ∈ s) (hxa : x ≠ a) (hya : y ≠ a) (hxy : x ≤ y) : + (f x - f a) / (x - a) ≤ (f y - f a) / (y - a) := +begin + rcases eq_or_lt_of_le hxy with rfl | hxy, + { simp }, + cases lt_or_gt_of_ne hxa with hxa hxa, + { cases lt_or_gt_of_ne hya with hya hya, + { convert hf.secant_mono_aux3 hx ha hxy hya using 1; + rw ← neg_div_neg_eq; + field_simp, }, + { convert hf.slope_mono_adjacent hx hy hxa hya using 1, + rw ← neg_div_neg_eq; + field_simp, } }, + { exact hf.secant_mono_aux2 ha hy hxa hxy, }, +end + +lemma strict_convex_on.secant_strict_mono_aux1 (hf : strict_convex_on 𝕜 s f) + {x y z : 𝕜} (hx : x ∈ s) (hz : z ∈ s) (hxy : x < y) (hyz : y < z) : + (z - x) * f y < (z - y) * f x + (y - x) * f z := +begin + have hxy' : 0 < y - x := by linarith, + have hyz' : 0 < z - y := by linarith, + have hxz' : 0 < z - x := by linarith, + rw ← lt_div_iff' hxz', + have ha : 0 < (z - y) / (z - x) := by positivity, + have hb : 0 < (y - x) / (z - x) := by positivity, + calc f y = f ((z - y) / (z - x) * x + (y - x) / (z - x) * z) : _ + ... < (z - y) / (z - x) * f x + (y - x) / (z - x) * f z : hf.2 hx hz (by linarith) ha hb _ + ... = ((z - y) * f x + (y - x) * f z) / (z - x) : _, + { congr' 1, + field_simp [hxy'.ne', hyz'.ne', hxz'.ne'], + ring }, + { field_simp [hxy'.ne', hyz'.ne', hxz'.ne'] }, + { field_simp [hxy'.ne', hyz'.ne', hxz'.ne'] } +end + +lemma strict_convex_on.secant_strict_mono_aux2 (hf : strict_convex_on 𝕜 s f) + {x y z : 𝕜} (hx : x ∈ s) (hz : z ∈ s) (hxy : x < y) (hyz : y < z) : + (f y - f x) / (y - x) < (f z - f x) / (z - x) := +begin + have hxy' : 0 < y - x := by linarith, + have hxz' : 0 < z - x := by linarith, + rw div_lt_div_iff hxy' hxz', + linarith only [hf.secant_strict_mono_aux1 hx hz hxy hyz], +end + +lemma strict_convex_on.secant_strict_mono_aux3 (hf : strict_convex_on 𝕜 s f) + {x y z : 𝕜} (hx : x ∈ s) (hz : z ∈ s) (hxy : x < y) (hyz : y < z) : + (f z - f x) / (z - x) < (f z - f y) / (z - y) := +begin + have hyz' : 0 < z - y := by linarith, + have hxz' : 0 < z - x := by linarith, + rw div_lt_div_iff hxz' hyz', + linarith only [hf.secant_strict_mono_aux1 hx hz hxy hyz], +end + +lemma strict_convex_on.secant_strict_mono (hf : strict_convex_on 𝕜 s f) + {a x y : 𝕜} (ha : a ∈ s) (hx : x ∈ s) (hy : y ∈ s) (hxa : x ≠ a) (hya : y ≠ a) (hxy : x < y) : + (f x - f a) / (x - a) < (f y - f a) / (y - a) := +begin + cases lt_or_gt_of_ne hxa with hxa hxa, + { cases lt_or_gt_of_ne hya with hya hya, + { convert hf.secant_strict_mono_aux3 hx ha hxy hya using 1; + rw ← neg_div_neg_eq; + field_simp, }, + { convert hf.slope_strict_mono_adjacent hx hy hxa hya using 1, + rw ← neg_div_neg_eq; + field_simp, } }, + { exact hf.secant_strict_mono_aux2 ha hy hxa hxy, }, +end + +lemma strict_concave_on.secant_strict_mono (hf : strict_concave_on 𝕜 s f) + {a x y : 𝕜} (ha : a ∈ s) (hx : x ∈ s) (hy : y ∈ s) (hxa : x ≠ a) (hya : y ≠ a) (hxy : x < y) : + (f y - f a) / (y - a) < (f x - f a) / (x - a) := +begin + have key := hf.neg.secant_strict_mono ha hx hy hxa hya hxy, + simp only [pi.neg_apply] at key, + rw ← neg_lt_neg_iff, + convert key using 1; field_simp, +end + /-- If `f` is convex on a set `s` in a linearly ordered field, and `f x < f y` for two points `x < y` in `s`, then `f` is strictly monotone on `s ∩ [y, ∞)`. -/ lemma convex_on.strict_mono_of_lt (hf : convex_on 𝕜 s f) diff --git a/src/analysis/convex/specific_functions.lean b/src/analysis/convex/specific_functions.lean index f6e81716af065..5c9eae9377903 100644 --- a/src/analysis/convex/specific_functions.lean +++ b/src/analysis/convex/specific_functions.lean @@ -5,6 +5,7 @@ Authors: Yury Kudryashov, Sébastien Gouëzel -/ import analysis.special_functions.pow.deriv import analysis.special_functions.sqrt +import tactic.linear_combination /-! # Collection of convex functions @@ -32,22 +33,86 @@ For `p : ℝ`, prove that `λ x, x ^ p` is concave when `0 ≤ p ≤ 1` and stri open real set open_locale big_operators nnreal -/-- `exp` is strictly convex on the whole real line. -/ +/-- `exp` is strictly convex on the whole real line. + +We give an elementary proof rather than using the second derivative test, since this lemma is +needed early in the analysis library. -/ lemma strict_convex_on_exp : strict_convex_on ℝ univ exp := -strict_convex_on_univ_of_deriv2_pos continuous_exp (λ x, (iter_deriv_exp 2).symm ▸ exp_pos x) +begin + apply strict_convex_on_of_slope_strict_mono_adjacent convex_univ, + rintros x y z - - hxy hyz, + transitivity exp y, + { have h1 : 0 < y - x := by linarith, + have h2 : x - y < 0 := by linarith, + rw div_lt_iff h1, + calc exp y - exp x = exp y - exp y * exp (x - y) : by rw ← exp_add; ring_nf + ... = exp y * (1 - exp (x - y)) : by ring + ... < exp y * (-(x - y)) : mul_lt_mul_of_pos_left _ y.exp_pos + ... = exp y * (y - x) : by ring, + linarith [add_one_lt_exp_of_nonzero h2.ne] }, + { have h1 : 0 < z - y := by linarith, + rw lt_div_iff h1, + calc exp y * (z - y) < exp y * (exp (z - y) - 1) : mul_lt_mul_of_pos_left _ y.exp_pos + ... = exp (z - y) * exp y - exp y : by ring + ... ≤ exp z - exp y : by rw ← exp_add; ring_nf, + linarith [add_one_lt_exp_of_nonzero h1.ne'] }, +end /-- `exp` is convex on the whole real line. -/ lemma convex_on_exp : convex_on ℝ univ exp := strict_convex_on_exp.convex_on -/-- `x^n`, `n : ℕ` is convex on the whole real line whenever `n` is even -/ +/-- `x^n`, `n : ℕ` is convex on `[0, +∞)` for all `n`. + +We give an elementary proof rather than using the second derivative test, since this lemma is +needed early in the analysis library. -/ +lemma convex_on_pow (n : ℕ) : convex_on ℝ (Ici 0) (λ x : ℝ, x^n) := +begin + induction n with k IH, + { exact convex_on_const (1:ℝ) (convex_Ici _) }, + refine ⟨convex_Ici _, _⟩, + rintros a (ha : 0 ≤ a) b (hb : 0 ≤ b) μ ν hμ hν h, + have H := IH.2 ha hb hμ hν h, + have : 0 ≤ (b ^ k - a ^ k) * (b - a) * μ * ν, + { cases le_or_lt a b with hab hab, + { have : a ^ k ≤ b ^ k := pow_le_pow_of_le_left ha hab k, + have : 0 ≤ (b ^ k - a ^ k) * (b - a) := by nlinarith, + positivity, }, + { have : b ^ k ≤ a ^ k := pow_le_pow_of_le_left hb hab.le k, + have : 0 ≤ (b ^ k - a ^ k) * (b - a) := by nlinarith, + positivity, } }, + calc (μ * a + ν * b) ^ k.succ = (μ * a + ν * b) * (μ * a + ν * b) ^ k : by ring_exp + ... ≤ (μ * a + ν * b) * (μ * a ^ k + ν * b ^ k) : mul_le_mul_of_nonneg_left H (by positivity) + ... ≤ (μ * a + ν * b) * (μ * a ^ k + ν * b ^ k) + (b ^ k - a ^ k) * (b - a) * μ * ν : by linarith + ... = (μ + ν) * (μ * a ^ k.succ + ν * b ^ k.succ) : by ring_exp + ... = μ * a ^ k.succ + ν * b ^ k.succ : by rw h; ring_exp, +end + +/-- `x^n`, `n : ℕ` is strictly convex on `[0, +∞)` for all `n` greater than `2`. -/ +lemma strict_convex_on_pow {n : ℕ} (hn : 2 ≤ n) : strict_convex_on ℝ (Ici 0) (λ x : ℝ, x^n) := +begin + apply strict_mono_on.strict_convex_on_of_deriv (convex_Ici _) (continuous_on_pow _), + rw [deriv_pow', interior_Ici], + exact λ x (hx : 0 < x) y hy hxy, mul_lt_mul_of_pos_left (pow_lt_pow_of_lt_left hxy hx.le $ + nat.sub_pos_of_lt hn) (nat.cast_pos.2 $ zero_lt_two.trans_le hn), +end + +/-- `x^n`, `n : ℕ` is convex on the whole real line whenever `n` is even. + +We give an elementary proof rather than using the second derivative test, since this lemma is +needed early in the analysis library. -/ lemma even.convex_on_pow {n : ℕ} (hn : even n) : convex_on ℝ set.univ (λ x : ℝ, x^n) := begin - apply convex_on_univ_of_deriv2_nonneg (differentiable_pow n), - { simp only [deriv_pow', differentiable.mul, differentiable_const, differentiable_pow] }, - { intro x, - obtain ⟨k, hk⟩ := (hn.tsub $ even_bit0 _).exists_two_nsmul _, - rw [iter_deriv_pow, finset.prod_range_cast_nat_sub, hk, nsmul_eq_mul, pow_mul'], - exact mul_nonneg (nat.cast_nonneg _) (pow_two_nonneg _) } + refine ⟨convex_univ, _⟩, + intros a ha b hb μ ν hμ hν h, + obtain ⟨k, rfl⟩ := hn.exists_two_nsmul _, + have : 0 ≤ (a - b) ^ 2 * μ * ν := by positivity, + calc (μ * a + ν * b) ^ (2 * k) = ((μ * a + ν * b) ^ 2) ^ k : by rw pow_mul + ... ≤ ((μ + ν) * (μ * a ^ 2 + ν * b ^ 2)) ^ k : pow_le_pow_of_le_left (by positivity) _ k + ... = (μ * a ^ 2 + ν * b ^ 2) ^ k : by rw h; ring_exp + ... ≤ μ * (a ^ 2) ^ k + ν * (b ^ 2) ^ k : _ + ... ≤ μ * a ^ (2 * k) + ν * b ^ (2 * k) : by ring_exp, + { linarith }, + { refine (convex_on_pow k).2 _ _ hμ hν h; dsimp; positivity }, end /-- `x^n`, `n : ℕ` is strictly convex on the whole real line whenever `n ≠ 0` is even. -/ @@ -61,26 +126,6 @@ begin (nat.cast_pos.2 h), end -/-- `x^n`, `n : ℕ` is convex on `[0, +∞)` for all `n` -/ -lemma convex_on_pow (n : ℕ) : convex_on ℝ (Ici 0) (λ x : ℝ, x^n) := -begin - apply convex_on_of_deriv2_nonneg (convex_Ici _) (continuous_pow n).continuous_on - (differentiable_on_pow n), - { simp only [deriv_pow'], exact (@differentiable_on_pow ℝ _ _ _).const_mul (n : ℝ) }, - { intros x hx, - rw [iter_deriv_pow, finset.prod_range_cast_nat_sub], - exact mul_nonneg (nat.cast_nonneg _) (pow_nonneg (interior_subset hx) _) } -end - -/-- `x^n`, `n : ℕ` is strictly convex on `[0, +∞)` for all `n` greater than `2`. -/ -lemma strict_convex_on_pow {n : ℕ} (hn : 2 ≤ n) : strict_convex_on ℝ (Ici 0) (λ x : ℝ, x^n) := -begin - apply strict_mono_on.strict_convex_on_of_deriv (convex_Ici _) (continuous_on_pow _), - rw [deriv_pow', interior_Ici], - exact λ x (hx : 0 < x) y hy hxy, mul_lt_mul_of_pos_left (pow_lt_pow_of_lt_left hxy hx.le $ - nat.sub_pos_of_lt hn) (nat.cast_pos.2 $ zero_lt_two.trans_le hn), -end - /-- Specific case of Jensen's inequality for sums of powers -/ lemma real.pow_sum_div_card_le_sum_pow {α : Type*} {s : finset α} {f : α → ℝ} (n : ℕ) (hf : ∀ a ∈ s, 0 ≤ f a) : (∑ x in s, f x) ^ (n + 1) / s.card ^ n ≤ ∑ x in s, (f x) ^ (n + 1) := @@ -140,20 +185,38 @@ begin exact ⟨int.coe_zero_le _, int.coe_nat_lt.2 $ finset.mem_range.1 ha⟩, end -/-- `x^m`, `m : ℤ` is convex on `(0, +∞)` for all `m` -/ -lemma convex_on_zpow (m : ℤ) : convex_on ℝ (Ioi 0) (λ x : ℝ, x^m) := +/-- `x^m`, `m : ℤ` is convex on `(0, +∞)` for all `m`. + +We give an elementary proof rather than using the second derivative test, since this lemma is +needed early in the analysis library. -/ +lemma convex_on_zpow : ∀ m : ℤ, convex_on ℝ (Ioi 0) (λ x : ℝ, x^m) +| (n : ℕ) := begin - have : ∀ n : ℤ, differentiable_on ℝ (λ x, x ^ n) (Ioi (0 : ℝ)), - from λ n, differentiable_on_zpow _ _ (or.inl $ lt_irrefl _), - apply convex_on_of_deriv2_nonneg (convex_Ioi 0); - try { simp only [interior_Ioi, deriv_zpow'] }, - { exact (this _).continuous_on }, - { exact this _ }, - { exact (this _).const_mul _ }, - { intros x hx, - rw iter_deriv_zpow, - refine mul_nonneg _ (zpow_nonneg (le_of_lt hx) _), - exact_mod_cast int_prod_range_nonneg _ _ (even_bit0 1) } + simp_rw zpow_coe_nat, + exact (convex_on_pow n).subset Ioi_subset_Ici_self (convex_Ioi _) +end +| -[1+ n] := +begin + simp_rw zpow_neg_succ_of_nat, + refine ⟨convex_Ioi _, _⟩, + rintros a (ha : 0 < a) b (hb : 0 < b) μ ν hμ hν h, + have ha' : 0 < a ^ (n + 1) := by positivity, + have hb' : 0 < b ^ (n + 1) := by positivity, + field_simp [ha.ne', hb.ne', ha'.ne', hb'.ne'], + rw div_le_div_iff, + { calc 1 * (a ^ (n + 1) * b ^ (n + 1)) + = ((μ + ν) ^ 2 * (a * b)) ^ (n + 1) : by rw h; ring_exp + ... ≤ ((μ * b + ν * a) * (μ * a + ν * b)) ^ (n + 1) : pow_le_pow_of_le_left _ _ _ + ... = (μ * b + ν * a) ^ (n + 1) * (μ * a + ν * b) ^ (n + 1) : by rw mul_pow + ... ≤ (μ * b ^ (n + 1) + ν * a ^ (n + 1)) * (μ * a + ν * b) ^ (n + 1) : _, + { positivity }, + { have : 0 ≤ μ * ν * (a - b) ^ 2 := by positivity, + linarith }, + { apply mul_le_mul_of_nonneg_right ((convex_on_pow (n + 1)).2 hb.le ha.le hμ hν h), + positivity } }, + { have : 0 < μ * a + ν * b := by cases le_or_lt a b; nlinarith, + positivity }, + { positivity }, end /-- `x^m`, `m : ℤ` is convex on `(0, +∞)` for all `m` except `0` and `1`. -/ @@ -171,56 +234,143 @@ begin fin_cases hm; cc, end -lemma convex_on_rpow {p : ℝ} (hp : 1 ≤ p) : convex_on ℝ (Ici 0) (λ x : ℝ, x^p) := +/- `real.log` is strictly concave on $(0, +∞)$. + +We give an elementary proof rather than using the second derivative test, since this lemma is +needed early in the analysis library. -/ +lemma strict_concave_on_log_Ioi : strict_concave_on ℝ (Ioi 0) log := +begin + apply strict_concave_on_of_slope_strict_anti_adjacent (convex_Ioi (0:ℝ)), + rintros x y z (hx : 0 < x) (hz : 0 < z) hxy hyz, + have hy : 0 < y := hx.trans hxy, + transitivity y⁻¹, + { have h : 0 < z - y := by linarith, + rw div_lt_iff h, + have hyz' : 0 < z / y := by positivity, + have hyz'' : z / y ≠ 1, + { contrapose! h, + rw div_eq_one_iff_eq hy.ne' at h, + simp [h] }, + calc log z - log y = log (z / y) : by rw ← log_div hz.ne' hy.ne' + ... < z / y - 1 : log_lt_sub_one_of_pos hyz' hyz'' + ... = y⁻¹ * (z - y) : by field_simp [hy.ne'] }, + { have h : 0 < y - x := by linarith, + rw lt_div_iff h, + have hxy' : 0 < x / y := by positivity, + have hxy'' : x / y ≠ 1, + { contrapose! h, + rw div_eq_one_iff_eq hy.ne' at h, + simp [h] }, + calc y⁻¹ * (y - x) = 1 - x / y : by field_simp [hy.ne'] + ... < - log (x / y) : by linarith [log_lt_sub_one_of_pos hxy' hxy''] + ... = - (log x - log y) : by rw log_div hx.ne' hy.ne' + ... = log y - log x : by ring }, +end + +/-- **Bernoulli's inequality** for real exponents, strict version: for `1 < p` and `-1 ≤ s`, with +`s ≠ 0`, we have `1 + p * s < (1 + s) ^ p`. -/ +lemma one_add_mul_self_lt_rpow_one_add {s : ℝ} (hs : -1 ≤ s) (hs' : s ≠ 0) {p : ℝ} (hp : 1 < p) : + 1 + p * s < (1 + s) ^ p := begin - have A : deriv (λ (x : ℝ), x ^ p) = λ x, p * x^(p-1), by { ext x, simp [hp] }, - apply convex_on_of_deriv2_nonneg (convex_Ici 0), - { exact continuous_on_id.rpow_const (λ x _, or.inr (zero_le_one.trans hp)) }, - { exact (differentiable_rpow_const hp).differentiable_on }, - { rw A, - assume x hx, - replace hx : x ≠ 0, by { simp at hx, exact ne_of_gt hx }, - simp [differentiable_at.differentiable_within_at, hx] }, - { assume x hx, - replace hx : 0 < x, by simpa using hx, - suffices : 0 ≤ p * ((p - 1) * x ^ (p - 1 - 1)), by simpa [ne_of_gt hx, A], - apply mul_nonneg (le_trans zero_le_one hp), - exact mul_nonneg (sub_nonneg_of_le hp) (rpow_nonneg_of_nonneg hx.le _) } + rcases eq_or_lt_of_le hs with rfl | hs, + { have : p ≠ 0 := by positivity, + simpa [zero_rpow this], }, + have hs1 : 0 < 1 + s := by linarith, + cases le_or_lt (1 + p * s) 0 with hs2 hs2, + { exact hs2.trans_lt (rpow_pos_of_pos hs1 _) }, + rw [rpow_def_of_pos hs1, ← exp_log hs2], + apply exp_strict_mono, + have hp : 0 < p := by positivity, + have hs3 : 1 + s ≠ 1 := by contrapose! hs'; linarith, + have hs4 : 1 + p * s ≠ 1 := by contrapose! hs'; nlinarith, + cases lt_or_gt_of_ne hs' with hs' hs', + { rw [← div_lt_iff hp, ← div_lt_div_right_of_neg hs'], + convert strict_concave_on_log_Ioi.secant_strict_mono zero_lt_one hs2 hs1 hs4 hs3 _ using 1, + { field_simp [log_one] }, + { field_simp [log_one] }, + { nlinarith } }, + { rw [← div_lt_iff hp, ← div_lt_div_right hs'], + convert strict_concave_on_log_Ioi.secant_strict_mono zero_lt_one hs1 hs2 hs3 hs4 _ using 1, + { field_simp [log_one, hp.ne'], }, + { field_simp [log_one] }, + { nlinarith } }, end +/-- **Bernoulli's inequality** for real exponents, non-strict version: for `1 ≤ p` and `-1 ≤ s` +we have `1 + p * s ≤ (1 + s) ^ p`. -/ +lemma one_add_mul_self_le_rpow_one_add {s : ℝ} (hs : -1 ≤ s) {p : ℝ} (hp : 1 ≤ p) : + 1 + p * s ≤ (1 + s) ^ p := +begin + rcases eq_or_lt_of_le hp with rfl | hp, + { simp }, + by_cases hs' : s = 0, + { simp [hs'] }, + exact (one_add_mul_self_lt_rpow_one_add hs hs' hp).le, +end + +/- For `p : ℝ` with `1 < p`, `λ x, x ^ p` is strictly convex on $[0, +∞)$. + +We give an elementary proof rather than using the second derivative test, since this lemma is +needed early in the analysis library. -/ lemma strict_convex_on_rpow {p : ℝ} (hp : 1 < p) : strict_convex_on ℝ (Ici 0) (λ x : ℝ, x^p) := begin - have A : deriv (λ (x : ℝ), x ^ p) = λ x, p * x^(p-1), by { ext x, simp [hp.le] }, - apply strict_convex_on_of_deriv2_pos (convex_Ici 0), - { exact continuous_on_id.rpow_const (λ x _, or.inr (zero_le_one.trans hp.le)) }, - rw interior_Ici, - rintro x (hx : 0 < x), - suffices : 0 < p * ((p - 1) * x ^ (p - 1 - 1)), by simpa [ne_of_gt hx, A], - exact mul_pos (zero_lt_one.trans hp) (mul_pos (sub_pos_of_lt hp) (rpow_pos_of_pos hx _)), + apply strict_convex_on_of_slope_strict_mono_adjacent (convex_Ici (0:ℝ)), + rintros x y z (hx : 0 ≤ x) (hz : 0 ≤ z) hxy hyz, + have hy : 0 < y := by linarith, + have hy' : 0 < y ^ p := rpow_pos_of_pos hy _, + have H1 : y ^ ((p - 1) + 1) = y ^ (p - 1) * y := rpow_add_one hy.ne' _, + ring_nf at H1, + transitivity p * y ^ (p - 1), + { have hyx' : x - y < 0 := by linarith only [hxy], + have h3 : 0 < y - x := by linarith only [hxy], + have hyx'' : x / y < 1 := by rwa div_lt_one hy, + have hyx''' : x / y - 1 < 0 := by linarith only [hyx''], + have hyx'''' : 0 ≤ x / y := by positivity, + have hyx''''' : -1 ≤ x / y - 1 := by linarith only [hyx''''], + have : 1 - (1 + ((x / y) - 1)) ^ p < - p * ((x / y) - 1), + { linarith [one_add_mul_self_lt_rpow_one_add hyx''''' hyx'''.ne hp] }, + rw [div_lt_iff h3, ← div_lt_div_right hy'], + convert this using 1, + { have H : (x / y) ^ p = x ^ p / y ^ p := div_rpow hx hy.le _, + ring_nf at ⊢ H, + field_simp [hy.ne', hy'.ne'] at ⊢ H, + linear_combination H }, + { field_simp [hy.ne', hy'.ne'], + linear_combination p * (-y + x) * H1 }, }, + { have hyz' : 0 < z - y := by linarith only [hyz], + have hyz'' : 1 < z / y := by rwa one_lt_div hy, + have hyz''' : 0 < z / y - 1 := by linarith only [hyz''], + have hyz'''' : -1 ≤ z / y - 1 := by linarith only [hyz''], + have : p * ((z / y) - 1) < (1 + ((z / y) - 1)) ^ p - 1, + { linarith [one_add_mul_self_lt_rpow_one_add hyz'''' hyz'''.ne' hp] }, + rw [lt_div_iff hyz', ← div_lt_div_right hy'], + convert this using 1, + { field_simp [hy.ne', hy'.ne'], + linear_combination - p * (z - y) * H1, }, + { have H : (z / y) ^ p = z ^ p / y ^ p := div_rpow hz hy.le _, + ring_nf at ⊢ H, + field_simp [hy.ne', hy'.ne'] at ⊢ H, + linear_combination -H } }, end -lemma strict_concave_on_log_Ioi : strict_concave_on ℝ (Ioi 0) log := +lemma convex_on_rpow {p : ℝ} (hp : 1 ≤ p) : convex_on ℝ (Ici 0) (λ x : ℝ, x^p) := begin - have h₁ : Ioi 0 ⊆ ({0} : set ℝ)ᶜ, - { exact λ x (hx : 0 < x) (hx' : x = 0), hx.ne' hx' }, - refine strict_concave_on_of_deriv2_neg' (convex_Ioi 0) - (continuous_on_log.mono h₁) (λ x (hx : 0 < x), _), - rw [function.iterate_succ, function.iterate_one], - change (deriv (deriv log)) x < 0, - rw [deriv_log', deriv_inv], - exact neg_neg_of_pos (inv_pos.2 $ sq_pos_of_ne_zero _ hx.ne'), + rcases eq_or_lt_of_le hp with rfl | hp, + { simpa using convex_on_id (convex_Ici _), }, + exact (strict_convex_on_rpow hp).convex_on, end lemma strict_concave_on_log_Iio : strict_concave_on ℝ (Iio 0) log := begin - have h₁ : Iio 0 ⊆ ({0} : set ℝ)ᶜ, - { exact λ x (hx : x < 0) (hx' : x = 0), hx.ne hx' }, - refine strict_concave_on_of_deriv2_neg' (convex_Iio 0) - (continuous_on_log.mono h₁) (λ x (hx : x < 0), _), - rw [function.iterate_succ, function.iterate_one], - change (deriv (deriv log)) x < 0, - rw [deriv_log', deriv_inv], - exact neg_neg_of_pos (inv_pos.2 $ sq_pos_of_ne_zero _ hx.ne), + refine ⟨convex_Iio _, _⟩, + rintros x (hx : x < 0) y (hy : y < 0) hxy a b ha hb hab, + have hx' : 0 < -x := by linarith, + have hy' : 0 < -y := by linarith, + have hxy' : - x ≠ - y := by contrapose! hxy; linarith, + calc a • log x + b • log y = a • log (-x) + b • log (-y) : by simp_rw [log_neg_eq_log] + ... < log (a • (-x) + b • (-y)) : strict_concave_on_log_Ioi.2 hx' hy' hxy' ha hb hab + ... = log (- (a • x + b • y)) : by congr' 1; simp only [algebra.id.smul_eq_mul]; ring + ... = _ : by rw log_neg_eq_log, end section sqrt_mul_log diff --git a/src/analysis/special_functions/log/basic.lean b/src/analysis/special_functions/log/basic.lean index 60b4961efe531..1f0069f4f4472 100644 --- a/src/analysis/special_functions/log/basic.lean +++ b/src/analysis/special_functions/log/basic.lean @@ -169,6 +169,14 @@ end lemma log_inj_on_pos : set.inj_on log (set.Ioi 0) := strict_mono_on_log.inj_on +lemma log_lt_sub_one_of_pos (hx1 : 0 < x) (hx2 : x ≠ 1) : log x < x - 1 := +begin + have h : log x ≠ 0, + { rw [← log_one, log_inj_on_pos.ne_iff hx1 zero_lt_one], + exact hx2 }, + linarith [add_one_lt_exp_of_nonzero h, exp_log hx1], +end + lemma eq_one_of_pos_of_log_eq_zero {x : ℝ} (h₁ : 0 < x) (h₂ : log x = 0) : x = 1 := log_inj_on_pos (set.mem_Ioi.2 h₁) (set.mem_Ioi.2 zero_lt_one) (h₂.trans real.log_one.symm) diff --git a/src/data/complex/exponential.lean b/src/data/complex/exponential.lean index 9591e4896fee2..4bdb7cd7d4d5c 100644 --- a/src/data/complex/exponential.lean +++ b/src/data/complex/exponential.lean @@ -1159,23 +1159,33 @@ by rw ← of_real_inj; simp [sinh_three_mul] open is_absolute_value +lemma sum_le_exp_of_nonneg {x : ℝ} (hx : 0 ≤ x) (n : ℕ) : ∑ i in range n, x ^ i / i! ≤ exp x := +calc ∑ i in range n, x ^ i / i! ≤ lim (⟨_, is_cau_seq_re (exp' x)⟩ : cau_seq ℝ has_abs.abs) : + begin + refine le_lim (cau_seq.le_of_exists ⟨n, λ j hj, _⟩), + simp only [exp', const_apply, mk_to_fun, re_sum], + norm_cast, + rw [← nat.add_sub_of_le hj, finset.sum_range_add], + refine le_add_of_nonneg_right (sum_nonneg (λ i hi, _)), + positivity, + end +... = exp x : by rw [exp, complex.exp, ← cau_seq_re, lim_re] + +lemma quadratic_le_exp_of_nonneg {x : ℝ} (hx : 0 ≤ x) : 1 + x + x ^ 2 / 2 ≤ exp x := +calc 1 + x + x ^ 2 / 2 = ∑ i in range 3, x ^ i / i! : by simp [finset.sum_range_succ] +... ≤ exp x : sum_le_exp_of_nonneg hx 3 + +lemma add_one_lt_exp_of_pos {x : ℝ} (hx : 0 < x) : x + 1 < exp x := +(by nlinarith : x + 1 < 1 + x + x ^ 2 / 2).trans_le (quadratic_le_exp_of_nonneg hx.le) + /-- This is an intermediate result that is later replaced by `real.add_one_le_exp`; use that lemma instead. -/ lemma add_one_le_exp_of_nonneg {x : ℝ} (hx : 0 ≤ x) : x + 1 ≤ exp x := -calc x + 1 ≤ lim (⟨(λ n : ℕ, ((exp' x) n).re), is_cau_seq_re (exp' x)⟩ : cau_seq ℝ has_abs.abs) : - le_lim (cau_seq.le_of_exists ⟨2, - λ j hj, show x + (1 : ℝ) ≤ (∑ m in range j, (x ^ m / m! : ℂ)).re, - from have h₁ : (((λ m : ℕ, (x ^ m / m! : ℂ)) ∘ nat.succ) 0).re = x, by simp, - have h₂ : ((x : ℂ) ^ 0 / 0!).re = 1, by simp, - begin - rw [← tsub_add_cancel_of_le hj, sum_range_succ', sum_range_succ', - add_re, add_re, h₁, h₂, add_assoc, - ← coe_re_add_group_hom, (re_add_group_hom).map_sum, coe_re_add_group_hom ], - refine le_add_of_nonneg_of_le (sum_nonneg (λ m hm, _)) le_rfl, - rw [← of_real_pow, ← of_real_nat_cast, ← of_real_div, of_real_re], - exact div_nonneg (pow_nonneg hx _) (nat.cast_nonneg _), - end⟩) -... = exp x : by rw [exp, complex.exp, ← cau_seq_re, lim_re] +begin + rcases eq_or_lt_of_le hx with rfl | h, + { simp }, + exact (add_one_lt_exp_of_pos h).le +end lemma one_le_exp {x : ℝ} (hx : 0 ≤ x) : 1 ≤ exp x := by linarith [add_one_le_exp_of_nonneg hx] @@ -1561,71 +1571,57 @@ calc cos 2 = cos (2 * 1) : congr_arg cos (mul_one _).symm zero_le_two) _ ... < 0 : by norm_num -lemma exp_bound_div_one_sub_of_interval_approx {x : ℝ} (h1 : 0 ≤ x) (h2 : x ≤ 1) : - ∑ (j : ℕ) in finset.range 3, x ^ j / (j.factorial) - + x ^ 3 * ((3 : ℕ) + 1) / ((3 : ℕ).factorial * (3 : ℕ)) - ≤ ∑ j in (finset.range 3), x ^ j := -begin - norm_num [finset.sum], - rw [add_assoc, add_comm (x + 1) (x ^ 3 * 4 / 18), ← add_assoc, add_le_add_iff_right, - ← add_le_add_iff_left (-(x ^ 2 / 2)), ← add_assoc, comm_ring.add_left_neg (x ^ 2 / 2), - zero_add, neg_add_eq_sub, sub_half, sq, pow_succ, sq], - have i1 : x * 4 / 18 ≤ 1 / 2 := by linarith, - have i2 : 0 ≤ x * 4 / 18 := by linarith, - have i3 := mul_le_mul h1 h1 le_rfl h1, - rw zero_mul at i3, - have t := mul_le_mul le_rfl i1 i2 i3, - rw ← mul_assoc, - rwa [mul_one_div, ← mul_div_assoc, ← mul_assoc] at t, -end +lemma exp_bound_div_one_sub_of_interval' {x : ℝ} (h1 : 0 < x) (h2 : x < 1) : + real.exp x < 1 / (1 - x) := +have H : 0 < 1 - (1 + x + x ^ 2) * (1 - x) := +calc 0 < x ^ 3 : by positivity +... = 1 - (1 + x + x ^ 2) * (1 - x) : by ring, +calc exp x ≤ _ : exp_bound' h1.le h2.le zero_lt_three +... ≤ 1 + x + x ^ 2 : by norm_num [finset.sum]; nlinarith +... < 1 / (1 - x) : by rw lt_div_iff; nlinarith lemma exp_bound_div_one_sub_of_interval {x : ℝ} (h1 : 0 ≤ x) (h2 : x < 1) : real.exp x ≤ 1 / (1 - x) := begin - have h : ∑ j in (finset.range 3), x ^ j ≤ 1 / (1 - x), - { norm_num [finset.sum], - have h1x : 0 < 1 - x := by simpa, - rw le_div_iff h1x, - norm_num [← add_assoc, mul_sub_left_distrib, mul_one, add_mul, - sub_add_eq_sub_sub, pow_succ' x 2], - have hx3 : 0 ≤ x ^ 3, - { norm_num, - exact h1 }, - linarith }, - exact (exp_bound' h1 h2.le $ by linarith).trans - ((exp_bound_div_one_sub_of_interval_approx h1 h2.le).trans h), + rcases eq_or_lt_of_le h1 with rfl | h1, + { simp }, + { exact (exp_bound_div_one_sub_of_interval' h1 h2).le } +end + +lemma one_sub_lt_exp_minus_of_pos {y : ℝ} (h : 0 < y) : 1 - y < real.exp (-y) := +begin + cases le_or_lt 1 y with h' h', + { linarith [(-y).exp_pos] }, + rw [exp_neg, lt_inv _ y.exp_pos, inv_eq_one_div], + { exact exp_bound_div_one_sub_of_interval' h h' }, + { linarith }, +end + +lemma one_sub_le_exp_minus_of_nonneg {y : ℝ} (h : 0 ≤ y) : 1 - y ≤ real.exp (-y) := +begin + rcases eq_or_lt_of_le h with rfl | h, + { simp }, + { exact (one_sub_lt_exp_minus_of_pos h).le } end -lemma one_sub_le_exp_minus_of_pos {y : ℝ} (h : 0 ≤ y) : 1 - y ≤ real.exp (-y) := +lemma add_one_lt_exp_of_neg {x : ℝ} (h : x < 0) : x + 1 < real.exp x := begin - rw real.exp_neg, - have r1 : (1 - y) * (real.exp y) ≤ 1, - { cases le_or_lt (1 - y) 0, - { have h'' : (1 - y) * y.exp ≤ 0, - { rw mul_nonpos_iff, - right, - exact ⟨h_1, y.exp_pos.le⟩ }, - linarith }, - have hy1 : y < 1 := by linarith, - rw ← le_div_iff' h_1, - exact exp_bound_div_one_sub_of_interval h hy1 }, - rw inv_eq_one_div, - rw le_div_iff' y.exp_pos, - rwa mul_comm at r1, + have h1 : 0 < -x := by linarith, + simpa [add_comm] using one_sub_lt_exp_minus_of_pos h1 end -lemma add_one_le_exp_of_nonpos {x : ℝ} (h : x ≤ 0) : x + 1 ≤ real.exp x := +lemma add_one_lt_exp_of_nonzero {x : ℝ} (hx : x ≠ 0) : x + 1 < real.exp x := begin - rw add_comm, - have h1 : 0 ≤ -x := by linarith, - simpa using one_sub_le_exp_minus_of_pos h1 + cases lt_or_gt_of_ne hx, + { exact real.add_one_lt_exp_of_neg h }, + exact add_one_lt_exp_of_pos h, end lemma add_one_le_exp (x : ℝ) : x + 1 ≤ real.exp x := begin cases le_or_lt 0 x, { exact real.add_one_le_exp_of_nonneg h }, - exact add_one_le_exp_of_nonpos h.le, + exact (add_one_lt_exp_of_neg h).le, end lemma one_sub_div_pow_le_exp_neg {n : ℕ} {t : ℝ} (ht' : t ≤ n) : (1 - t / n) ^ n ≤ exp (-t) := From 7daeaf3072304c498b653628add84a88d0e78767 Mon Sep 17 00:00:00 2001 From: jakelev Date: Thu, 18 May 2023 05:56:34 +0000 Subject: [PATCH 1026/1291] feat(nat/factorial): add double factorials (#18994) Add double factorials. --- src/data/nat/factorial/double_factorial.lean | 72 ++++++++++++++++++++ 1 file changed, 72 insertions(+) create mode 100644 src/data/nat/factorial/double_factorial.lean diff --git a/src/data/nat/factorial/double_factorial.lean b/src/data/nat/factorial/double_factorial.lean new file mode 100644 index 0000000000000..87b621056a376 --- /dev/null +++ b/src/data/nat/factorial/double_factorial.lean @@ -0,0 +1,72 @@ +/- +Copyright (c) 2023 Jake Levinson. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jake Levinson +-/ +import data.nat.factorial.basic +import algebra.big_operators.order +import tactic.ring +/-! +# Double factorials + +This file defines the double factorial, + `n‼ := n * (n - 2) * (n - 4) * ...`. + +## Main declarations + +* `nat.double_factorial`: The double factorial. +-/ + +open_locale nat +namespace nat + +/-- `nat.double_factorial n` is the double factorial of `n`. -/ +@[simp] def double_factorial : ℕ → ℕ +| 0 := 1 +| 1 := 1 +| (k + 2) := (k + 2) * double_factorial k + +-- This notation is `\!!` not two !'s +localized "notation (name := nat.double_factorial) n `‼`:10000 := nat.double_factorial n" in nat + +lemma double_factorial_add_two (n : ℕ) : (n + 2)‼ = (n + 2) * n‼ := rfl + +lemma double_factorial_add_one (n : ℕ) : (n + 1)‼ = (n + 1) * (n - 1)‼ := by { cases n; refl } + +lemma factorial_eq_mul_double_factorial : ∀ (n : ℕ), (n + 1)! = (n + 1)‼ * n‼ +| 0 := rfl +| (k + 1) := begin + rw [double_factorial_add_two, factorial, factorial_eq_mul_double_factorial, + mul_comm _ (k‼), mul_assoc] +end + +lemma double_factorial_two_mul : + ∀ (n : ℕ), (2 * n)‼ = 2^n * n! +| 0 := rfl +| (n + 1) := begin + rw [mul_add, mul_one, double_factorial_add_two, factorial, pow_succ, + double_factorial_two_mul, succ_eq_add_one], + ring, +end + +open_locale big_operators + +lemma double_factorial_eq_prod_even : + ∀ (n : ℕ), (2 * n)‼ = ∏ i in finset.range n, (2 * (i + 1)) +| 0 := rfl +| (n + 1) := begin + rw [finset.prod_range_succ, ← double_factorial_eq_prod_even, mul_comm (2 * n)‼, + (by ring : 2 * (n + 1) = 2 * n + 2)], + refl, +end + +lemma double_factorial_eq_prod_odd : + ∀ (n : ℕ), (2 * n + 1)‼ = ∏ i in finset.range n, (2 * (i + 1) + 1) +| 0 := rfl +| (n + 1) := begin + rw [finset.prod_range_succ, ← double_factorial_eq_prod_odd, mul_comm (2 * n + 1)‼, + (by ring : 2 * (n + 1) + 1 = (2 * n + 1) + 2)], + refl, +end + +end nat From 56b71f0b55c03f70332b862e65c3aa1aa1249ca1 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Thu, 18 May 2023 07:01:15 +0000 Subject: [PATCH 1027/1291] feat(algebra/category/Module/change_of_rings): coextension of scalars (#15958) Co-authored-by: Jujian Zhang --- .../category/Module/change_of_rings.lean | 186 ++++++++++++++++++ 1 file changed, 186 insertions(+) diff --git a/src/algebra/category/Module/change_of_rings.lean b/src/algebra/category/Module/change_of_rings.lean index fcf032e3c20f0..b52900dc79917 100644 --- a/src/algebra/category/Module/change_of_rings.lean +++ b/src/algebra/category/Module/change_of_rings.lean @@ -20,10 +20,16 @@ import ring_theory.tensor_product module structure is defined by `s • (s' ⊗ m) := (s * s') ⊗ m` and `R`-linear map `l : M ⟶ M'` is sent to `S`-linear map `s ⊗ m ↦ s ⊗ l m : S ⨂ M ⟶ S ⨂ M'`. +* `category_theory.Module.coextend_scalars`: given rings `R, S` and a ring homomorphism `R ⟶ S` + then `coextend_scalars : Module R ⥤ Module S` is defined by `M ↦ (S →ₗ[R] M)` where `S` is seen as + `R-module` by restriction of scalars and `l ↦ l ∘ _`. + ## Main results * `category_theory.Module.extend_restrict_scalars_adj`: given commutative rings `R, S` and a ring homomorphism `f : R →+* S`, the extension and restriction of scalars by `f` are adjoint functors. +* `category_theory.Module.restrict_coextend_scalars_adj`: given rings `R, S` and a ring homomorphism + `f : R ⟶ S` then `coextend_scalars f` is the right adjoint of `restrict_scalars f`. ## List of notations Let `R, S` be rings and `f : R →+* S` @@ -170,6 +176,186 @@ variables {R : Type u₁} {S : Type u₂} [comm_ring R] [comm_ring S] (f : R → end extend_scalars +namespace coextend_scalars + +variables {R : Type u₁} {S : Type u₂} [ring R] [ring S] (f : R →+* S) + +section unbundled + +variables (M : Type v) [add_comm_monoid M] [module R M] + +-- We use `S'` to denote `S` viewed as `R`-module, via the map `f`. +local notation `S'` := (restrict_scalars f).obj ⟨S⟩ + +/-- + Given an `R`-module M, consider Hom(S, M) -- the `R`-linear maps between S (as an `R`-module by + means of restriction of scalars) and M. `S` acts on Hom(S, M) by `s • g = x ↦ g (x • s)` + -/ +instance has_smul : has_smul S $ S' →ₗ[R] M := +{ smul := λ s g, + { to_fun := λ (s' : S), g (s' * s : S), + map_add' := λ (x y : S), by simp [add_mul, map_add], + map_smul' := λ r (t : S), by rw [ring_hom.id_apply, @restrict_scalars.smul_def _ _ _ _ f ⟨S⟩, + ←linear_map.map_smul, @restrict_scalars.smul_def _ _ _ _ f ⟨S⟩, smul_eq_mul, smul_eq_mul, + mul_assoc] } } + +@[simp] lemma smul_apply' (s : S) (g : S' →ₗ[R] M) (s' : S) : + @has_smul.smul _ _ (coextend_scalars.has_smul f _) s g s' = g (s' * s : S) := rfl + +instance mul_action : mul_action S $ S' →ₗ[R] M := +{ one_smul := λ g, linear_map.ext $ λ (s : S), by simp, + mul_smul := λ (s t : S) g, linear_map.ext $ λ (x : S), by simp [mul_assoc], + ..coextend_scalars.has_smul f _ } + +instance distrib_mul_action : distrib_mul_action S $ S' →ₗ[R] M := +{ smul_add := λ s g h, linear_map.ext $ λ (t : S), by simp, + smul_zero := λ s, linear_map.ext $ λ (t : S), by simp, + ..coextend_scalars.mul_action f _ } + +/-- +`S` acts on Hom(S, M) by `s • g = x ↦ g (x • s)`, this action defines an `S`-module structure on +Hom(S, M). + -/ +instance is_module : module S $ S' →ₗ[R] M := +{ add_smul := λ s1 s2 g, linear_map.ext $ λ (x : S), by simp [mul_add], + zero_smul := λ g, linear_map.ext $ λ (x : S), by simp, + ..coextend_scalars.distrib_mul_action f _ } + +end unbundled + +variable (M : Module.{v} R) + +/-- If `M` is an `R`-module, then the set of `R`-linear maps `S →ₗ[R] M` is an `S`-module with +scalar multiplication defined by `s • l := x ↦ l (x • s)`-/ +def obj' : Module S := ⟨(restrict_scalars f).obj ⟨S⟩ →ₗ[R] M⟩ + +instance : has_coe_to_fun (obj' f M) (λ g, S → M) := +{ coe := λ g, g.to_fun } + +/-- If `M, M'` are `R`-modules, then any `R`-linear map `g : M ⟶ M'` induces an `S`-linear map +`(S →ₗ[R] M) ⟶ (S →ₗ[R] M')` defined by `h ↦ g ∘ h`-/ +@[simps] def map' {M M' : Module R} (g : M ⟶ M') : obj' f M ⟶ obj' f M' := +{ to_fun := λ h, g.comp h, + map_add' := λ _ _, linear_map.comp_add _ _ _, + map_smul' := λ s h, linear_map.ext $ λ (t : S), by simpa only [smul_apply'] } + +end coextend_scalars + +/-- +For any rings `R, S` and a ring homomorphism `f : R →+* S`, there is a functor from `R`-module to +`S`-module defined by `M ↦ (S →ₗ[R] M)` where `S` is considered as an `R`-module via restriction of +scalars and `g : M ⟶ M'` is sent to `h ↦ g ∘ h`. +-/ +def coextend_scalars {R : Type u₁} {S : Type u₂} [ring R] [ring S] (f : R →+* S) : + Module R ⥤ Module S := +{ obj := coextend_scalars.obj' f, + map := λ _ _, coextend_scalars.map' f, + map_id' := λ M, linear_map.ext $ λ h, linear_map.ext $ λ x, rfl, + map_comp' := λ _ _ _ g h, linear_map.ext $ λ h, linear_map.ext $ λ x, rfl } + +namespace coextend_scalars + +variables {R : Type u₁} {S : Type u₂} [ring R] [ring S] (f : R →+* S) + +instance (M : Module R) : has_coe_to_fun ((coextend_scalars f).obj M) (λ g, S → M) := +(infer_instance : has_coe_to_fun (coextend_scalars.obj' f M) _) + +lemma smul_apply (M : Module R) (g : (coextend_scalars f).obj M) (s s' : S) : + (s • g) s' = g (s' * s) := rfl + +@[simp] lemma map_apply {M M' : Module R} (g : M ⟶ M') (x) (s : S) : + (coextend_scalars f).map g x s = g (x s) := rfl + +end coextend_scalars + +namespace restriction_coextension_adj + +variables {R : Type u₁} {S : Type u₂} [ring R] [ring S] (f : R →+* S) + +/-- +Given `R`-module X and `S`-module Y, any `g : (restrict_of_scalars f).obj Y ⟶ X` +corresponds to `Y ⟶ (coextend_scalars f).obj X` by sending `y ↦ (s ↦ g (s • y))` +-/ +@[simps] def hom_equiv.from_restriction {X Y} (g : (restrict_scalars f).obj Y ⟶ X) : + Y ⟶ (coextend_scalars f).obj X := +{ to_fun := λ (y : Y), + { to_fun := λ (s : S), g $ (s • y : Y), + map_add' := λ (s1 s2 : S), by simp [add_smul], + map_smul' := λ r (s : S), by rw [ring_hom.id_apply, ←g.map_smul, + @restrict_scalars.smul_def _ _ _ _ f ⟨S⟩, smul_eq_mul, mul_smul, + @restrict_scalars.smul_def _ _ _ _ f Y] }, + map_add' := λ (y1 y2 : Y), linear_map.ext $ λ (s : S), + by rw [linear_map.add_apply, linear_map.coe_mk, linear_map.coe_mk, linear_map.coe_mk, + smul_add, map_add], + map_smul' := λ s y, linear_map.ext $ λ (t : S), by simp [mul_smul] } + +/-- +Given `R`-module X and `S`-module Y, any `g : Y ⟶ (coextend_scalars f).obj X` +corresponds to `(restrict_scalars f).obj Y ⟶ X` by `y ↦ g y 1` +-/ +@[simps] def hom_equiv.to_restriction {X Y} (g : Y ⟶ (coextend_scalars f).obj X) : + (restrict_scalars f).obj Y ⟶ X := +{ to_fun := λ (y : Y), (g y).to_fun (1 : S), + map_add' := λ x y, by simp only [g.map_add, linear_map.to_fun_eq_coe, linear_map.add_apply], + map_smul' := λ r (y : Y), by rw [linear_map.to_fun_eq_coe, linear_map.to_fun_eq_coe, + ring_hom.id_apply, ←linear_map.map_smul, restrict_scalars.smul_def f r y, + @restrict_scalars.smul_def _ _ _ _ f ⟨S⟩, smul_eq_mul, mul_one, linear_map.map_smul, + coextend_scalars.smul_apply, one_mul], } + +/-- +The natural transformation from identity functor to the composition of restriction and coextension +of scalars. +-/ +@[simps] protected def unit' : 𝟭 (Module S) ⟶ restrict_scalars f ⋙ coextend_scalars f := +{ app := λ Y, + { to_fun := λ (y : Y), + { to_fun := λ (s : S), (s • y : Y), + map_add' := λ s s', add_smul _ _ _, + map_smul' := λ r (s : S), by rw [ring_hom.id_apply, @restrict_scalars.smul_def _ _ _ _ f ⟨S⟩, + smul_eq_mul, mul_smul, restrict_scalars.smul_def f] }, + map_add' := λ y1 y2, linear_map.ext $ λ (s : S), by rw [linear_map.add_apply, linear_map.coe_mk, + linear_map.coe_mk, linear_map.coe_mk, smul_add], + map_smul' := λ s (y : Y), linear_map.ext $ λ (t : S), by simp [mul_smul] }, + naturality' := λ Y Y' g, linear_map.ext $ λ (y : Y), linear_map.ext $ λ (s : S), + by simp [coextend_scalars.map_apply] } + +/-- +The natural transformation from the composition of coextension and restriction of scalars to +identity functor. +-/ +@[simps] protected def counit' : coextend_scalars f ⋙ restrict_scalars f ⟶ 𝟭 (Module R) := +{ app := λ X, + { to_fun := λ g, g.to_fun (1 : S), + map_add' := λ x1 x2, by simp [linear_map.to_fun_eq_coe], + map_smul' := λ r (g : (restrict_scalars f).obj ((coextend_scalars f).obj X)), + begin + simp only [linear_map.to_fun_eq_coe, ring_hom.id_apply], + rw [restrict_scalars.smul_def f, coextend_scalars.smul_apply, one_mul, ←linear_map.map_smul, + @restrict_scalars.smul_def _ _ _ _ f ⟨S⟩, smul_eq_mul, mul_one], + end }, + naturality' := λ X X' g, linear_map.ext $ λ h, by simp [coextend_scalars.map_apply] } + +end restriction_coextension_adj + +/-- Restriction of scalars is left adjoint to coextension of scalars. -/ +@[simps] def restrict_coextend_scalars_adj {R : Type u₁} {S : Type u₂} [ring R] [ring S] + (f : R →+* S) : restrict_scalars f ⊣ coextend_scalars f := +{ hom_equiv := λ X Y, + { to_fun := restriction_coextension_adj.hom_equiv.from_restriction f, + inv_fun := restriction_coextension_adj.hom_equiv.to_restriction f, + left_inv := λ g, linear_map.ext $ λ (x : X), by simp, + right_inv := λ g, linear_map.ext $ λ x, linear_map.ext $ λ (s : S), by simp }, + unit := restriction_coextension_adj.unit' f, + counit := restriction_coextension_adj.counit' f, + hom_equiv_unit' := λ X Y g, linear_map.ext $ λ y, rfl, + hom_equiv_counit' := λ Y X g, linear_map.ext $ λ (y : Y), by simp } + +instance {R : Type u₁} {S : Type u₂} [ring R] [ring S] (f : R →+* S) : + category_theory.is_left_adjoint (restrict_scalars f) := ⟨_, restrict_coextend_scalars_adj f⟩ + +instance {R : Type u₁} {S : Type u₂} [ring R] [ring S] (f : R →+* S) : + category_theory.is_right_adjoint (coextend_scalars f) := ⟨_, restrict_coextend_scalars_adj f⟩ + namespace extend_restrict_scalars_adj open_locale change_of_rings From c720ca1664115159ac610a74e079287d052cf8d0 Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Thu, 18 May 2023 07:01:16 +0000 Subject: [PATCH 1028/1291] chore(number_theory/modular_forms): don't define jacobi_theta on subtype (#19029) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit In a previous PR, the Jacobi theta function was defined on the subtype `ℍ` of `ℂ` and this is slightly annoying to work with. This PR tweaks it to be a function on `ℂ` (whose values outside `ℍ` are uninteresting, but well-defined). Split off from #19027. --- .../modular_forms/jacobi_theta.lean | 82 ++++++++++--------- 1 file changed, 44 insertions(+), 38 deletions(-) diff --git a/src/number_theory/modular_forms/jacobi_theta.lean b/src/number_theory/modular_forms/jacobi_theta.lean index dd6578c6a37ae..1afe11de17ba4 100644 --- a/src/number_theory/modular_forms/jacobi_theta.lean +++ b/src/number_theory/modular_forms/jacobi_theta.lean @@ -24,7 +24,7 @@ open complex real asymptotics open_locale real big_operators upper_half_plane manifold /-- Jacobi's theta function `∑' (n : ℤ), exp (π * I * n ^ 2 * τ)`. -/ -noncomputable def jacobi_theta (τ : ℍ) : ℂ := ∑' (n : ℤ), cexp (π * I * n ^ 2 * τ) +noncomputable def jacobi_theta (z : ℂ) : ℂ := ∑' (n : ℤ), cexp (π * I * n ^ 2 * z) lemma norm_exp_mul_sq_le {z : ℂ} (hz : 0 < z.im) (n : ℤ) : ‖cexp (π * I * n ^ 2 * z)‖ ≤ exp (-π * z.im) ^ n.nat_abs := @@ -63,23 +63,26 @@ lemma summable_exp_mul_sq {z : ℂ} (hz : 0 < z.im) : let ⟨bd, h, h'⟩ := exists_summable_bound_exp_mul_sq hz in summable_norm_iff.mp (summable_of_nonneg_of_le (λ n, norm_nonneg _) (h' $ le_refl _) h) -lemma jacobi_theta_two_vadd (τ : ℍ) : jacobi_theta ((2 : ℝ) +ᵥ τ) = jacobi_theta τ := +lemma jacobi_theta_two_add (z : ℂ) : jacobi_theta (2 + z) = jacobi_theta z := begin refine tsum_congr (λ n, _), - rw [upper_half_plane.coe_vadd, of_real_bit0, of_real_one], suffices : cexp (↑π * I * ↑n ^ 2 * 2) = 1, by rw [mul_add, complex.exp_add, this, one_mul], rw [(by { push_cast, ring } : ↑π * I * ↑n ^ 2 * 2 = ↑(n ^ 2) * (2 * π * I)), complex.exp_int_mul, complex.exp_two_pi_mul_I, one_zpow], end -lemma jacobi_theta_T_sq_smul (τ : ℍ) : jacobi_theta (modular_group.T ^ 2 • τ) = jacobi_theta τ := +lemma jacobi_theta_T_sq_smul (τ : ℍ) : + jacobi_theta ↑(modular_group.T ^ 2 • τ) = jacobi_theta τ := begin - suffices : (2 : ℝ) +ᵥ τ = modular_group.T ^ (2 : ℤ) • τ, from this ▸ (jacobi_theta_two_vadd τ), - simp only [←subtype.coe_inj, upper_half_plane.modular_T_zpow_smul, int.cast_two], + suffices : ↑(modular_group.T ^ 2 • τ) = (2 : ℂ) + ↑τ, + { simp_rw [this, jacobi_theta_two_add] }, + have : modular_group.T ^ (2 : ℕ) = modular_group.T ^ (2 : ℤ), by refl, + simp_rw [this, upper_half_plane.modular_T_zpow_smul, upper_half_plane.coe_vadd], + push_cast, end lemma jacobi_theta_S_smul (τ : ℍ) : - jacobi_theta (modular_group.S • τ) = (-I * τ) ^ (1 / 2 : ℂ) * jacobi_theta τ := + jacobi_theta ↑(modular_group.S • τ) = (-I * τ) ^ (1 / 2 : ℂ) * jacobi_theta τ := begin unfold jacobi_theta, rw [upper_half_plane.modular_S_smul, upper_half_plane.coe_mk], @@ -104,10 +107,10 @@ begin ring_nf } end -lemma has_sum_nat_jacobi_theta (τ : ℍ) : - has_sum (λ (n : ℕ), cexp (π * I * (n + 1) ^ 2 * τ)) ((jacobi_theta τ - 1) / 2) := +lemma has_sum_nat_jacobi_theta {z : ℂ} (hz : 0 < im z) : + has_sum (λ (n : ℕ), cexp (π * I * (n + 1) ^ 2 * z)) ((jacobi_theta z - 1) / 2) := begin - have := (summable_exp_mul_sq τ.im_pos).has_sum.sum_nat_of_sum_int, + have := (summable_exp_mul_sq hz).has_sum.sum_nat_of_sum_int, rw ←@has_sum_nat_add_iff' ℂ _ _ _ _ 1 at this, simp_rw [finset.sum_range_one, int.cast_neg, int.cast_coe_nat, nat.cast_zero, neg_zero, int.cast_zero, sq (0:ℂ), mul_zero, zero_mul, neg_sq, ←mul_two, complex.exp_zero, @@ -117,30 +120,30 @@ begin simp_rw mul_div_cancel _ two_ne_zero, end -lemma jacobi_theta_eq_tsum_nat (τ : ℍ) : - jacobi_theta τ = 1 + 2 * ∑' (n : ℕ), cexp (π * I * (n + 1) ^ 2 * τ) := -by rw [(has_sum_nat_jacobi_theta τ).tsum_eq, mul_div_cancel' _ (two_ne_zero' ℂ), ←add_sub_assoc, +lemma jacobi_theta_eq_tsum_nat {z : ℂ} (hz : 0 < im z) : + jacobi_theta z = 1 + 2 * ∑' (n : ℕ), cexp (π * I * (n + 1) ^ 2 * z) := +by rw [(has_sum_nat_jacobi_theta hz).tsum_eq, mul_div_cancel' _ (two_ne_zero' ℂ), ←add_sub_assoc, add_sub_cancel'] /-- An explicit upper bound for `‖jacobi_theta τ - 1‖`. -/ -lemma norm_jacobi_theta_sub_one_le (τ : ℍ) : - ‖jacobi_theta τ - 1‖ ≤ 2 / (1 - exp (-π * τ.im)) * exp (-π * τ.im) := +lemma norm_jacobi_theta_sub_one_le {z : ℂ} (hz : 0 < im z) : + ‖jacobi_theta z - 1‖ ≤ 2 / (1 - exp (-π * z.im)) * exp (-π * z.im) := begin - suffices : ‖∑' (n : ℕ), cexp (π * I * (n + 1) ^ 2 * τ)‖ ≤ exp (-π * τ.im) / (1 - exp (-π * τ.im)), - { calc ‖jacobi_theta τ - 1‖ = 2 * ‖∑' (n : ℕ), cexp (π * I * (n + 1) ^ 2 * τ)‖ : - by rw [sub_eq_iff_eq_add'.mpr (jacobi_theta_eq_tsum_nat τ), norm_mul, complex.norm_eq_abs, + suffices : ‖∑' (n : ℕ), cexp (π * I * (n + 1) ^ 2 * z)‖ ≤ exp (-π * z.im) / (1 - exp (-π * z.im)), + { calc ‖jacobi_theta z - 1‖ = 2 * ‖∑' (n : ℕ), cexp (π * I * (n + 1) ^ 2 * z)‖ : + by rw [sub_eq_iff_eq_add'.mpr (jacobi_theta_eq_tsum_nat hz), norm_mul, complex.norm_eq_abs, complex.abs_two] - ... ≤ 2 * (rexp (-π * τ.im) / (1 - rexp (-π * τ.im))) : + ... ≤ 2 * (rexp (-π * z.im) / (1 - rexp (-π * z.im))) : by rwa [mul_le_mul_left (zero_lt_two' ℝ)] - ... = 2 / (1 - rexp (-π * τ.im)) * rexp (-π * τ.im) : by rw [div_mul_comm, mul_comm] }, - have : ∀ (n : ℕ), ‖cexp (π * I * (n + 1) ^ 2 * τ)‖ ≤ exp (-π * τ.im) ^ (n + 1), + ... = 2 / (1 - rexp (-π * z.im)) * rexp (-π * z.im) : by rw [div_mul_comm, mul_comm] }, + have : ∀ (n : ℕ), ‖cexp (π * I * (n + 1) ^ 2 * z)‖ ≤ exp (-π * z.im) ^ (n + 1), { intro n, - simpa only [int.cast_add, int.cast_one] using norm_exp_mul_sq_le τ.im_pos (n + 1) }, - have s : has_sum (λ n : ℕ, rexp (-π * τ.im) ^ (n + 1)) (exp (-π * τ.im) / (1 - exp (-π * τ.im))), + simpa only [int.cast_add, int.cast_one] using norm_exp_mul_sq_le hz (n + 1) }, + have s : has_sum (λ n : ℕ, rexp (-π * z.im) ^ (n + 1)) (exp (-π * z.im) / (1 - exp (-π * z.im))), { simp_rw [pow_succ, div_eq_mul_inv, has_sum_mul_left_iff (real.exp_ne_zero _)], - exact has_sum_geometric_of_lt_1 (exp_pos (-π * τ.im)).le - (exp_lt_one_iff.mpr $ (mul_neg_of_neg_of_pos (neg_lt_zero.mpr pi_pos) τ.im_pos)) }, - have aux : summable (λ (n : ℕ), ‖cexp (↑π * I * (↑n + 1) ^ 2 * ↑τ)‖), + exact has_sum_geometric_of_lt_1 (exp_pos (-π * z.im)).le + (exp_lt_one_iff.mpr $ (mul_neg_of_neg_of_pos (neg_lt_zero.mpr pi_pos) hz)) }, + have aux : summable (λ (n : ℕ), ‖cexp (↑π * I * (↑n + 1) ^ 2 * z)‖), from summable_of_nonneg_of_le (λ n, norm_nonneg _) this s.summable, exact (norm_tsum_le_tsum_norm aux).trans ((tsum_mono aux s.summable this).trans (le_of_eq s.tsum_eq)), @@ -148,25 +151,27 @@ end /-- The norm of `jacobi_theta τ - 1` decays exponentially as `im τ → ∞`. -/ lemma is_O_at_im_infty_jacobi_theta_sub_one : - is_O upper_half_plane.at_im_infty (λ τ, jacobi_theta τ - 1) (λ τ, rexp (-π * τ.im)) := + is_O (filter.comap im filter.at_top) (λ τ, jacobi_theta τ - 1) (λ τ, rexp (-π * τ.im)) := begin - simp_rw [is_O, is_O_with, filter.eventually, upper_half_plane.at_im_infty_mem], - refine ⟨2 / (1 - rexp (-π)), 1, (λ τ hτ, (norm_jacobi_theta_sub_one_le τ).trans _)⟩, + simp_rw [is_O, is_O_with, filter.eventually_comap, filter.eventually_at_top], + refine ⟨2 / (1 - rexp (-π)), 1, λ y hy z hz, (norm_jacobi_theta_sub_one_le + (hz.symm ▸ (zero_lt_one.trans_le hy) : 0 < im z)).trans _⟩, rw [real.norm_eq_abs, real.abs_exp], refine mul_le_mul_of_nonneg_right _ (exp_pos _).le, rw [div_le_div_left (zero_lt_two' ℝ), sub_le_sub_iff_left, exp_le_exp, neg_mul, neg_le_neg_iff], - { exact le_mul_of_one_le_right pi_pos.le hτ }, - { rw [sub_pos, exp_lt_one_iff, neg_mul, neg_lt_zero], exact mul_pos pi_pos τ.im_pos }, + { exact le_mul_of_one_le_right pi_pos.le (hz.symm ▸ hy) }, + { rw [sub_pos, exp_lt_one_iff, neg_mul, neg_lt_zero], + exact mul_pos pi_pos (hz.symm ▸ (zero_lt_one.trans_le hy)) }, { rw [sub_pos, exp_lt_one_iff, neg_lt_zero], exact pi_pos } end -lemma differentiable_at_tsum_exp_mul_sq (τ : ℍ) : - differentiable_at ℂ (λ z, ∑' (n : ℤ), cexp (π * I * n ^ 2 * z)) ↑τ := +lemma differentiable_at_jacobi_theta {z : ℂ} (hz : 0 < im z) : + differentiable_at ℂ jacobi_theta z := begin suffices : ∀ (y : ℝ) (hy : 0 < y), differentiable_on ℂ (λ z, ∑' (n : ℤ), cexp (π * I * n ^ 2 * z)) {w : ℂ | y < im w}, - from let ⟨y, hy, hy'⟩ := exists_between τ.im_pos in (this y hy).differentiable_at - ((complex.continuous_im.is_open_preimage _ is_open_Ioi).mem_nhds (τ.coe_im ▸ hy')), + from let ⟨y, hy, hy'⟩ := exists_between hz in (this y hy).differentiable_at + ((complex.continuous_im.is_open_preimage _ is_open_Ioi).mem_nhds hy'), intros y hy, have h1 : ∀ (n : ℤ) (w : ℂ) (hw : y < im w), differentiable_within_at ℂ (λ (v : ℂ), cexp (↑π * I * ↑n ^ 2 * v)) {z : ℂ | y < im z} w, @@ -176,7 +181,8 @@ begin exact differentiable_on_tsum_of_summable_norm bd_s h1 h2 (λ i w hw, le_bd (le_of_lt hw) i), end -lemma mdifferentiable_jacobi_theta : mdifferentiable 𝓘(ℂ) 𝓘(ℂ) jacobi_theta := -λ τ, (differentiable_at_tsum_exp_mul_sq τ).mdifferentiable_at.comp τ τ.mdifferentiable_coe +lemma mdifferentiable_jacobi_theta : mdifferentiable 𝓘(ℂ) 𝓘(ℂ) (jacobi_theta ∘ coe : ℍ → ℂ) := +λ τ, (differentiable_at_jacobi_theta τ.2).mdifferentiable_at.comp τ τ.mdifferentiable_coe -lemma continuous_jacobi_theta : continuous jacobi_theta := mdifferentiable_jacobi_theta.continuous +lemma continuous_at_jacobi_theta {z : ℂ} (hz : 0 < im z) : + continuous_at jacobi_theta z := (differentiable_at_jacobi_theta hz).continuous_at From 50251fd6309cca5ca2e747882ffecd2729f38c5d Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Thu, 18 May 2023 07:01:17 +0000 Subject: [PATCH 1029/1291] chore(*): add mathlib4 synchronization comments (#19032) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.direct_sum.algebra` * `algebra.direct_sum.ring` * `algebra.homology.opposite` * `algebra.homology.quasi_iso` * `algebraic_topology.fundamental_groupoid.basic` * `analysis.box_integral.partition.subbox_induction` * `analysis.complex.circle` * `analysis.normed_space.is_R_or_C` * `analysis.normed_space.operator_norm` * `analysis.special_functions.trigonometric.angle` * `analysis.special_functions.trigonometric.basic` * `analysis.special_functions.trigonometric.inverse` * `category_theory.limits.constructions.over.products` * `category_theory.preadditive.endo_functor` * `combinatorics.simple_graph.ends.properties` * `combinatorics.simple_graph.finsubgraph` * `linear_algebra.multilinear.finite_dimensional` * `ring_theory.free_comm_ring` * `ring_theory.polynomial.quotient` * `ring_theory.tensor_product` * `topology.algebra.module.finite_dimension` * `topology.metric_space.partition_of_unity` --- src/algebra/direct_sum/algebra.lean | 3 +++ src/algebra/direct_sum/ring.lean | 3 +++ src/algebra/homology/opposite.lean | 3 +++ src/algebra/homology/quasi_iso.lean | 3 +++ src/algebraic_topology/fundamental_groupoid/basic.lean | 3 +++ src/analysis/box_integral/partition/subbox_induction.lean | 3 +++ src/analysis/complex/circle.lean | 3 +++ src/analysis/normed_space/is_R_or_C.lean | 3 +++ src/analysis/normed_space/operator_norm.lean | 3 +++ src/analysis/special_functions/trigonometric/angle.lean | 3 +++ src/analysis/special_functions/trigonometric/basic.lean | 3 +++ src/analysis/special_functions/trigonometric/inverse.lean | 3 +++ src/category_theory/limits/constructions/over/products.lean | 3 +++ src/category_theory/preadditive/endo_functor.lean | 3 +++ src/combinatorics/simple_graph/ends/properties.lean | 3 +++ src/combinatorics/simple_graph/finsubgraph.lean | 3 +++ src/linear_algebra/multilinear/finite_dimensional.lean | 3 +++ src/ring_theory/free_comm_ring.lean | 3 +++ src/ring_theory/polynomial/quotient.lean | 3 +++ src/ring_theory/tensor_product.lean | 3 +++ src/topology/algebra/module/finite_dimension.lean | 3 +++ src/topology/metric_space/partition_of_unity.lean | 3 +++ 22 files changed, 66 insertions(+) diff --git a/src/algebra/direct_sum/algebra.lean b/src/algebra/direct_sum/algebra.lean index 1e58cc310470a..e08cc480379f0 100644 --- a/src/algebra/direct_sum/algebra.lean +++ b/src/algebra/direct_sum/algebra.lean @@ -9,6 +9,9 @@ import algebra.direct_sum.ring /-! # Additively-graded algebra structures on `⨁ i, A i` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file provides `R`-algebra structures on external direct sums of `R`-modules. Recall that if `A i` are a family of `add_comm_monoid`s indexed by an `add_monoid`, then an instance diff --git a/src/algebra/direct_sum/ring.lean b/src/algebra/direct_sum/ring.lean index 8da08a69f5213..cca2fd311b35e 100644 --- a/src/algebra/direct_sum/ring.lean +++ b/src/algebra/direct_sum/ring.lean @@ -9,6 +9,9 @@ import algebra.direct_sum.basic /-! # Additively-graded multiplicative structures on `⨁ i, A i` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This module provides a set of heterogeneous typeclasses for defining a multiplicative structure over `⨁ i, A i` such that `(*) : A i → A j → A (i + j)`; that is to say, `A` forms an additively-graded ring. The typeclasses are: diff --git a/src/algebra/homology/opposite.lean b/src/algebra/homology/opposite.lean index e015eb1f61d5e..7613e58ab6a0d 100644 --- a/src/algebra/homology/opposite.lean +++ b/src/algebra/homology/opposite.lean @@ -10,6 +10,9 @@ import algebra.homology.additive /-! # Opposite categories of complexes + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. Given a preadditive category `V`, the opposite of its category of chain complexes is equivalent to the category of cochain complexes of objects in `Vᵒᵖ`. We define this equivalence, and another analagous equivalence (for a general category of homological complexes with a general diff --git a/src/algebra/homology/quasi_iso.lean b/src/algebra/homology/quasi_iso.lean index 240c744648cef..82fd875676257 100644 --- a/src/algebra/homology/quasi_iso.lean +++ b/src/algebra/homology/quasi_iso.lean @@ -9,6 +9,9 @@ import category_theory.abelian.homology /-! # Quasi-isomorphisms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A chain map is a quasi-isomorphism if it induces isomorphisms on homology. ## Future work diff --git a/src/algebraic_topology/fundamental_groupoid/basic.lean b/src/algebraic_topology/fundamental_groupoid/basic.lean index 9c3fea201bd58..90509950dcf73 100644 --- a/src/algebraic_topology/fundamental_groupoid/basic.lean +++ b/src/algebraic_topology/fundamental_groupoid/basic.lean @@ -11,6 +11,9 @@ import topology.homotopy.path /-! # Fundamental groupoid of a space +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given a topological space `X`, we can define the fundamental groupoid of `X` to be the category with objects being points of `X`, and morphisms `x ⟶ y` being paths from `x` to `y`, quotiented by homotopy equivalence. With this, the fundamental group of `X` based at `x` is just the automorphism diff --git a/src/analysis/box_integral/partition/subbox_induction.lean b/src/analysis/box_integral/partition/subbox_induction.lean index 230dbddcfe4a3..b059a5dae3ffc 100644 --- a/src/analysis/box_integral/partition/subbox_induction.lean +++ b/src/analysis/box_integral/partition/subbox_induction.lean @@ -9,6 +9,9 @@ import analysis.box_integral.partition.tagged /-! # Induction on subboxes +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove (see `box_integral.tagged_partition.exists_is_Henstock_is_subordinate_homothetic`) that for every box `I` in `ℝⁿ` and a function `r : ℝⁿ → ℝ` positive on `I` there exists a tagged partition `π` of `I` such diff --git a/src/analysis/complex/circle.lean b/src/analysis/complex/circle.lean index 66aebc42a9920..1be99911891dd 100644 --- a/src/analysis/complex/circle.lean +++ b/src/analysis/complex/circle.lean @@ -10,6 +10,9 @@ import analysis.normed.field.unit_ball /-! # The circle +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines `circle` to be the metric sphere (`metric.sphere`) in `ℂ` centred at `0` of radius `1`. We equip it with the following structure: diff --git a/src/analysis/normed_space/is_R_or_C.lean b/src/analysis/normed_space/is_R_or_C.lean index 0b19875c5996e..403f56e8030ab 100644 --- a/src/analysis/normed_space/is_R_or_C.lean +++ b/src/analysis/normed_space/is_R_or_C.lean @@ -10,6 +10,9 @@ import analysis.normed_space.pointwise /-! # Normed spaces over R or C +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file is about results on normed spaces over the fields `ℝ` and `ℂ`. ## Main definitions diff --git a/src/analysis/normed_space/operator_norm.lean b/src/analysis/normed_space/operator_norm.lean index a2a0e71b7316e..f497c31dbaba2 100644 --- a/src/analysis/normed_space/operator_norm.lean +++ b/src/analysis/normed_space/operator_norm.lean @@ -12,6 +12,9 @@ import topology.algebra.module.strong_topology /-! # Operator norm on the space of continuous linear maps +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Define the operator norm on the space of continuous (semi)linear maps between normed spaces, and prove its basic properties. In particular, show that this space is itself a normed space. diff --git a/src/analysis/special_functions/trigonometric/angle.lean b/src/analysis/special_functions/trigonometric/angle.lean index 422c30266f8b2..ef263b8b5c0ad 100644 --- a/src/analysis/special_functions/trigonometric/angle.lean +++ b/src/analysis/special_functions/trigonometric/angle.lean @@ -11,6 +11,9 @@ import topology.instances.sign /-! # The type of angles +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define `real.angle` to be the quotient group `ℝ/2πℤ` and prove a few simple lemmas about trigonometric functions and angles. -/ diff --git a/src/analysis/special_functions/trigonometric/basic.lean b/src/analysis/special_functions/trigonometric/basic.lean index 2473599ce5a67..63146c0dedc68 100644 --- a/src/analysis/special_functions/trigonometric/basic.lean +++ b/src/analysis/special_functions/trigonometric/basic.lean @@ -8,6 +8,9 @@ import analysis.special_functions.exp /-! # Trigonometric functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions This file contains the definition of `π`. diff --git a/src/analysis/special_functions/trigonometric/inverse.lean b/src/analysis/special_functions/trigonometric/inverse.lean index 81ce4c01ad38c..0d6c733ba7551 100644 --- a/src/analysis/special_functions/trigonometric/inverse.lean +++ b/src/analysis/special_functions/trigonometric/inverse.lean @@ -10,6 +10,9 @@ import topology.algebra.order.proj_Icc /-! # Inverse trigonometric functions. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + See also `analysis.special_functions.trigonometric.arctan` for the inverse tan function. (This is delayed as it is easier to set up after developing complex trigonometric functions.) diff --git a/src/category_theory/limits/constructions/over/products.lean b/src/category_theory/limits/constructions/over/products.lean index cfd8ca68c6ec8..d6f9073a3ab12 100644 --- a/src/category_theory/limits/constructions/over/products.lean +++ b/src/category_theory/limits/constructions/over/products.lean @@ -11,6 +11,9 @@ import category_theory.limits.shapes.finite_products /-! # Products in the over category +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Shows that products in the over category can be derived from wide pullbacks in the base category. The main result is `over_product_of_wide_pullback`, which says that if `C` has `J`-indexed wide pullbacks, then `over B` has `J`-indexed products. diff --git a/src/category_theory/preadditive/endo_functor.lean b/src/category_theory/preadditive/endo_functor.lean index c4c77a465cf8d..9bdf2f5e3cbf7 100644 --- a/src/category_theory/preadditive/endo_functor.lean +++ b/src/category_theory/preadditive/endo_functor.lean @@ -11,6 +11,9 @@ import category_theory.preadditive.additive_functor /-! # Preadditive structure on algebras over a monad +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + If `C` is a preadditive categories and `F` is an additive endofunctor on `C` then `algebra F` is also preadditive. Dually, the category `coalgebra F` is also preadditive. -/ diff --git a/src/combinatorics/simple_graph/ends/properties.lean b/src/combinatorics/simple_graph/ends/properties.lean index 1daaae2a569e2..009e5a47a1e97 100644 --- a/src/combinatorics/simple_graph/ends/properties.lean +++ b/src/combinatorics/simple_graph/ends/properties.lean @@ -7,6 +7,9 @@ import combinatorics.simple_graph.ends.defs /-! # Properties of the ends of graphs +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file is meant to contain results about the ends of (locally finite connected) graphs. -/ diff --git a/src/combinatorics/simple_graph/finsubgraph.lean b/src/combinatorics/simple_graph/finsubgraph.lean index 3d37c06ebae0a..dc13b8d50a0a5 100644 --- a/src/combinatorics/simple_graph/finsubgraph.lean +++ b/src/combinatorics/simple_graph/finsubgraph.lean @@ -9,6 +9,9 @@ import combinatorics.simple_graph.subgraph /-! # Homomorphisms from finite subgraphs +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the type of finite subgraphs of a `simple_graph` and proves a compactness result for homomorphisms to a finite codomain. diff --git a/src/linear_algebra/multilinear/finite_dimensional.lean b/src/linear_algebra/multilinear/finite_dimensional.lean index eba4f84553061..aa328791e6c38 100644 --- a/src/linear_algebra/multilinear/finite_dimensional.lean +++ b/src/linear_algebra/multilinear/finite_dimensional.lean @@ -8,6 +8,9 @@ import linear_algebra.free_module.finite.matrix /-! # Multilinear maps over finite dimensional spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The main results are that multilinear maps over finitely-generated, free modules are finitely-generated and free. diff --git a/src/ring_theory/free_comm_ring.lean b/src/ring_theory/free_comm_ring.lean index 887264d82f83e..7a3211dfad72a 100644 --- a/src/ring_theory/free_comm_ring.lean +++ b/src/ring_theory/free_comm_ring.lean @@ -11,6 +11,9 @@ import ring_theory.free_ring /-! # Free commutative rings +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The theory of the free commutative ring generated by a type `α`. It is isomorphic to the polynomial ring over ℤ with variables in `α` diff --git a/src/ring_theory/polynomial/quotient.lean b/src/ring_theory/polynomial/quotient.lean index f7cdff17c5fc2..08bb9a5c7ff89 100644 --- a/src/ring_theory/polynomial/quotient.lean +++ b/src/ring_theory/polynomial/quotient.lean @@ -10,6 +10,9 @@ import ring_theory.ideal.quotient_operations /-! # Quotients of polynomial rings + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ open_locale polynomial diff --git a/src/ring_theory/tensor_product.lean b/src/ring_theory/tensor_product.lean index 203263beb8cd1..75ebf694c0df5 100644 --- a/src/ring_theory/tensor_product.lean +++ b/src/ring_theory/tensor_product.lean @@ -11,6 +11,9 @@ import linear_algebra.direct_sum.finsupp /-! # The tensor product of R-algebras +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Let `R` be a (semi)ring and `A` an `R`-algebra. In this file we: diff --git a/src/topology/algebra/module/finite_dimension.lean b/src/topology/algebra/module/finite_dimension.lean index e5af4572eaa6c..fe65f4c3628e8 100644 --- a/src/topology/algebra/module/finite_dimension.lean +++ b/src/topology/algebra/module/finite_dimension.lean @@ -11,6 +11,9 @@ import topology.algebra.module.determinant /-! # Finite dimensional topological vector spaces over complete fields +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Let `𝕜` be a complete nontrivially normed field, and `E` a topological vector space (TVS) over `𝕜` (i.e we have `[add_comm_group E] [module 𝕜 E] [topological_space E] [topological_add_group E]` and `[has_continuous_smul 𝕜 E]`). diff --git a/src/topology/metric_space/partition_of_unity.lean b/src/topology/metric_space/partition_of_unity.lean index 9b79f18b41305..7088b4bb42ef6 100644 --- a/src/topology/metric_space/partition_of_unity.lean +++ b/src/topology/metric_space/partition_of_unity.lean @@ -9,6 +9,9 @@ import analysis.convex.partition_of_unity /-! # Lemmas about (e)metric spaces that need partition of unity +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The main lemma in this file (see `metric.exists_continuous_real_forall_closed_ball_subset`) says the following. Let `X` be a metric space. Let `K : ι → set X` be a locally finite family of closed sets, let `U : ι → set X` be a family of open sets such that `K i ⊆ U i` for all `i`. Then there exists a From df0098f0db291900600f32070f6abb3e178be2ba Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 18 May 2023 09:13:10 +0000 Subject: [PATCH 1030/1291] feat(field_theory/minpoly/basic): `minpoly` is unaffected by injective `alg_hom`s (#19030) --- src/field_theory/minpoly/basic.lean | 14 ++++++++++++-- 1 file changed, 12 insertions(+), 2 deletions(-) diff --git a/src/field_theory/minpoly/basic.lean b/src/field_theory/minpoly/basic.lean index 1538c6bad695a..121a0388520f2 100644 --- a/src/field_theory/minpoly/basic.lean +++ b/src/field_theory/minpoly/basic.lean @@ -17,7 +17,7 @@ such as ireducibility under the assumption `B` is a domain. open_locale classical polynomial open polynomial set function -variables {A B : Type*} +variables {A B B' : Type*} section min_poly_def variables (A) [comm_ring A] [ring B] [algebra A B] @@ -40,7 +40,7 @@ end min_poly_def namespace minpoly section ring -variables [comm_ring A] [ring B] [algebra A B] +variables [comm_ring A] [ring B] [ring B'] [algebra A B] [algebra A B'] variables {x : B} /-- A minimal polynomial is monic. -/ @@ -54,6 +54,16 @@ lemma ne_zero [nontrivial A] (hx : is_integral A x) : minpoly A x ≠ 0 := lemma eq_zero (hx : ¬ is_integral A x) : minpoly A x = 0 := dif_neg hx +lemma minpoly_alg_hom (f : B →ₐ[A] B') (hf : function.injective f) (x : B) : + minpoly A (f x) = minpoly A x := +begin + refine dif_ctx_congr (is_integral_alg_hom_iff _ hf) (λ _, _) (λ _, rfl), + simp_rw [←polynomial.aeval_def, aeval_alg_hom, alg_hom.comp_apply, _root_.map_eq_zero_iff f hf], +end + +@[simp] lemma minpoly_alg_equiv (f : B ≃ₐ[A] B') (x : B) : minpoly A (f x) = minpoly A x := +minpoly_alg_hom (f : B →ₐ[A] B') f.injective x + variables (A x) /-- An element is a root of its minimal polynomial. -/ From 8f9fea08977f7e450770933ee6abb20733b47c92 Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Thu, 18 May 2023 13:54:34 +0000 Subject: [PATCH 1031/1291] refactor(analysis/convex/specific_functions): split (#19031) Split `analysis/convex/specific_function` into a part which doesn't require differentiation (as of #19026) and a part which does. This removes the dependence of `measure_theory/integral/bochner` on differentiation, and decreases by 11 the length of the longest path in mathlib. See [Zulip](https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/Shortcut.20to.20integration) Co-authored-by: Scott Morrison --- .../calculus/bump_function_inner.lean | 2 + src/analysis/calculus/monotone.lean | 1 + .../calculus/parametric_integral.lean | 1 + .../basic.lean} | 187 ++---------------- .../convex/specific_functions/deriv.lean | 171 ++++++++++++++++ src/analysis/fourier/add_circle.lean | 2 +- src/analysis/fourier/fourier_transform.lean | 1 + src/analysis/mean_inequalities.lean | 4 +- src/analysis/mean_inequalities_pow.lean | 26 ++- src/analysis/normed_space/lp_space.lean | 1 + src/analysis/normed_space/pi_Lp.lean | 1 + .../special_functions/polar_coord.lean | 1 + .../trigonometric/complex.lean | 2 +- src/geometry/manifold/instances/sphere.lean | 2 +- src/measure_theory/function/jacobian.lean | 1 + src/measure_theory/function/lp_space.lean | 1 + src/measure_theory/integral/bochner.lean | 2 + .../integral/divergence_theorem.lean | 1 + .../integral/integral_eq_improper.lean | 1 + src/number_theory/bertrand.lean | 2 +- .../metric_space/hausdorff_dimension.lean | 1 + 21 files changed, 232 insertions(+), 179 deletions(-) rename src/analysis/convex/{specific_functions.lean => specific_functions/basic.lean} (59%) create mode 100644 src/analysis/convex/specific_functions/deriv.lean diff --git a/src/analysis/calculus/bump_function_inner.lean b/src/analysis/calculus/bump_function_inner.lean index f236cbd0c3333..2c1173af84735 100644 --- a/src/analysis/calculus/bump_function_inner.lean +++ b/src/analysis/calculus/bump_function_inner.lean @@ -3,8 +3,10 @@ Copyright (c) 2020 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel, Floris van Doorn -/ +import analysis.calculus.extend_deriv import analysis.calculus.iterated_deriv import analysis.inner_product_space.calculus +import analysis.special_functions.exp_deriv import measure_theory.integral.set_integral /-! diff --git a/src/analysis/calculus/monotone.lean b/src/analysis/calculus/monotone.lean index baf24c83fbf35..c94163209a778 100644 --- a/src/analysis/calculus/monotone.lean +++ b/src/analysis/calculus/monotone.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ +import analysis.calculus.deriv import measure_theory.covering.one_dim import order.monotone.extension diff --git a/src/analysis/calculus/parametric_integral.lean b/src/analysis/calculus/parametric_integral.lean index bdbb1d29fc49f..27b7e59312900 100644 --- a/src/analysis/calculus/parametric_integral.lean +++ b/src/analysis/calculus/parametric_integral.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Patrick Massot. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Patrick Massot -/ +import analysis.calculus.mean_value import measure_theory.integral.set_integral /-! diff --git a/src/analysis/convex/specific_functions.lean b/src/analysis/convex/specific_functions/basic.lean similarity index 59% rename from src/analysis/convex/specific_functions.lean rename to src/analysis/convex/specific_functions/basic.lean index 5c9eae9377903..a4ee577d8966b 100644 --- a/src/analysis/convex/specific_functions.lean +++ b/src/analysis/convex/specific_functions/basic.lean @@ -1,28 +1,30 @@ /- Copyright (c) 2020 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury Kudryashov, Sébastien Gouëzel +Authors: Yury Kudryashov, Sébastien Gouëzel, Heather Macbeth -/ -import analysis.special_functions.pow.deriv -import analysis.special_functions.sqrt +import analysis.convex.slope +import analysis.special_functions.pow.real import tactic.linear_combination /-! # Collection of convex functions -In this file we prove that the following functions are convex: +In this file we prove that the following functions are convex or strictly convex: * `strict_convex_on_exp` : The exponential function is strictly convex. -* `even.convex_on_pow`, `even.strict_convex_on_pow` : For an even `n : ℕ`, `λ x, x ^ n` is convex - and strictly convex when `2 ≤ n`. -* `convex_on_pow`, `strict_convex_on_pow` : For `n : ℕ`, `λ x, x ^ n` is convex on $[0, +∞)$ and - strictly convex when `2 ≤ n`. -* `convex_on_zpow`, `strict_convex_on_zpow` : For `m : ℤ`, `λ x, x ^ m` is convex on $[0, +∞)$ and - strictly convex when `m ≠ 0, 1`. -* `convex_on_rpow`, `strict_convex_on_rpow` : For `p : ℝ`, `λ x, x ^ p` is convex on $[0, +∞)$ when - `1 ≤ p` and strictly convex when `1 < p`. +* `even.convex_on_pow`: For an even `n : ℕ`, `λ x, x ^ n` is convex. +* `convex_on_pow`: For `n : ℕ`, `λ x, x ^ n` is convex on $[0, +∞)$. +* `convex_on_zpow`: For `m : ℤ`, `λ x, x ^ m` is convex on $[0, +∞)$. * `strict_concave_on_log_Ioi`, `strict_concave_on_log_Iio`: `real.log` is strictly concave on $(0, +∞)$ and $(-∞, 0)$ respectively. +* `convex_on_rpow`, `strict_convex_on_rpow` : For `p : ℝ`, `λ x, x ^ p` is convex on $[0, +∞)$ when + `1 ≤ p` and strictly convex when `1 < p`. + +The proofs in this file are deliberately elementary, *not* by appealing to the sign of the second +derivative. This is in order to keep this file early in the import hierarchy, since it is on the +path to Hölder's and Minkowski's inequalities and after that to Lp spaces and most of measure +theory. ## TODO @@ -87,15 +89,6 @@ begin ... = μ * a ^ k.succ + ν * b ^ k.succ : by rw h; ring_exp, end -/-- `x^n`, `n : ℕ` is strictly convex on `[0, +∞)` for all `n` greater than `2`. -/ -lemma strict_convex_on_pow {n : ℕ} (hn : 2 ≤ n) : strict_convex_on ℝ (Ici 0) (λ x : ℝ, x^n) := -begin - apply strict_mono_on.strict_convex_on_of_deriv (convex_Ici _) (continuous_on_pow _), - rw [deriv_pow', interior_Ici], - exact λ x (hx : 0 < x) y hy hxy, mul_lt_mul_of_pos_left (pow_lt_pow_of_lt_left hxy hx.le $ - nat.sub_pos_of_lt hn) (nat.cast_pos.2 $ zero_lt_two.trans_le hn), -end - /-- `x^n`, `n : ℕ` is convex on the whole real line whenever `n` is even. We give an elementary proof rather than using the second derivative test, since this lemma is @@ -115,76 +108,6 @@ begin { refine (convex_on_pow k).2 _ _ hμ hν h; dsimp; positivity }, end -/-- `x^n`, `n : ℕ` is strictly convex on the whole real line whenever `n ≠ 0` is even. -/ -lemma even.strict_convex_on_pow {n : ℕ} (hn : even n) (h : n ≠ 0) : - strict_convex_on ℝ set.univ (λ x : ℝ, x^n) := -begin - apply strict_mono.strict_convex_on_univ_of_deriv (continuous_pow n), - rw deriv_pow', - replace h := nat.pos_of_ne_zero h, - exact strict_mono.const_mul (odd.strict_mono_pow $ nat.even.sub_odd h hn $ nat.odd_iff.2 rfl) - (nat.cast_pos.2 h), -end - -/-- Specific case of Jensen's inequality for sums of powers -/ -lemma real.pow_sum_div_card_le_sum_pow {α : Type*} {s : finset α} {f : α → ℝ} (n : ℕ) - (hf : ∀ a ∈ s, 0 ≤ f a) : (∑ x in s, f x) ^ (n + 1) / s.card ^ n ≤ ∑ x in s, (f x) ^ (n + 1) := -begin - rcases s.eq_empty_or_nonempty with rfl | hs, - { simp_rw [finset.sum_empty, zero_pow' _ (nat.succ_ne_zero n), zero_div] }, - { have hs0 : 0 < (s.card : ℝ) := nat.cast_pos.2 hs.card_pos, - suffices : (∑ x in s, f x / s.card) ^ (n + 1) ≤ ∑ x in s, (f x ^ (n + 1) / s.card), - { rwa [← finset.sum_div, ← finset.sum_div, div_pow, pow_succ' (s.card : ℝ), - ← div_div, div_le_iff hs0, div_mul, div_self hs0.ne', div_one] at this }, - have := @convex_on.map_sum_le ℝ ℝ ℝ α _ _ _ _ _ _ (set.Ici 0) (λ x, x ^ (n + 1)) s - (λ _, 1 / s.card) (coe ∘ f) (convex_on_pow (n + 1)) _ _ (λ i hi, set.mem_Ici.2 (hf i hi)), - { simpa only [inv_mul_eq_div, one_div, algebra.id.smul_eq_mul] using this }, - { simp only [one_div, inv_nonneg, nat.cast_nonneg, implies_true_iff] }, - { simpa only [one_div, finset.sum_const, nsmul_eq_mul] using mul_inv_cancel hs0.ne' } } -end - -lemma nnreal.pow_sum_div_card_le_sum_pow {α : Type*} (s : finset α) (f : α → ℝ≥0) (n : ℕ) : - (∑ x in s, f x) ^ (n + 1) / s.card ^ n ≤ ∑ x in s, (f x) ^ (n + 1) := -by simpa only [← nnreal.coe_le_coe, nnreal.coe_sum, nonneg.coe_div, nnreal.coe_pow] using - @real.pow_sum_div_card_le_sum_pow α s (coe ∘ f) n (λ _ _, nnreal.coe_nonneg _) - -lemma finset.prod_nonneg_of_card_nonpos_even - {α β : Type*} [linear_ordered_comm_ring β] - {f : α → β} [decidable_pred (λ x, f x ≤ 0)] - {s : finset α} (h0 : even (s.filter (λ x, f x ≤ 0)).card) : - 0 ≤ ∏ x in s, f x := -calc 0 ≤ (∏ x in s, ((if f x ≤ 0 then (-1:β) else 1) * f x)) : - finset.prod_nonneg (λ x _, by - { split_ifs with hx hx, by simp [hx], simp at hx ⊢, exact le_of_lt hx }) -... = _ : by rw [finset.prod_mul_distrib, finset.prod_ite, finset.prod_const_one, - mul_one, finset.prod_const, neg_one_pow_eq_pow_mod_two, nat.even_iff.1 h0, pow_zero, one_mul] - -lemma int_prod_range_nonneg (m : ℤ) (n : ℕ) (hn : even n) : - 0 ≤ ∏ k in finset.range n, (m - k) := -begin - rcases hn with ⟨n, rfl⟩, - induction n with n ihn, { simp }, - rw ← two_mul at ihn, - rw [← two_mul, nat.succ_eq_add_one, mul_add, mul_one, bit0, ← add_assoc, finset.prod_range_succ, - finset.prod_range_succ, mul_assoc], - refine mul_nonneg ihn _, generalize : (1 + 1) * n = k, - cases le_or_lt m k with hmk hmk, - { have : m ≤ k + 1, from hmk.trans (lt_add_one ↑k).le, - convert mul_nonneg_of_nonpos_of_nonpos (sub_nonpos_of_le hmk) _, - convert sub_nonpos_of_le this }, - { exact mul_nonneg (sub_nonneg_of_le hmk.le) (sub_nonneg_of_le hmk) } -end - -lemma int_prod_range_pos {m : ℤ} {n : ℕ} (hn : even n) (hm : m ∉ Ico (0 : ℤ) n) : - 0 < ∏ k in finset.range n, (m - k) := -begin - refine (int_prod_range_nonneg m n hn).lt_of_ne (λ h, hm _), - rw [eq_comm, finset.prod_eq_zero_iff] at h, - obtain ⟨a, ha, h⟩ := h, - rw sub_eq_zero.1 h, - exact ⟨int.coe_zero_le _, int.coe_nat_lt.2 $ finset.mem_range.1 ha⟩, -end - /-- `x^m`, `m : ℤ` is convex on `(0, +∞)` for all `m`. We give an elementary proof rather than using the second derivative test, since this lemma is @@ -219,21 +142,6 @@ begin { positivity }, end -/-- `x^m`, `m : ℤ` is convex on `(0, +∞)` for all `m` except `0` and `1`. -/ -lemma strict_convex_on_zpow {m : ℤ} (hm₀ : m ≠ 0) (hm₁ : m ≠ 1) : - strict_convex_on ℝ (Ioi 0) (λ x : ℝ, x^m) := -begin - apply strict_convex_on_of_deriv2_pos' (convex_Ioi 0), - { exact (continuous_on_zpow₀ m).mono (λ x hx, ne_of_gt hx) }, - intros x hx, - rw iter_deriv_zpow, - refine mul_pos _ (zpow_pos_of_pos hx _), - exact_mod_cast int_prod_range_pos (even_bit0 1) (λ hm, _), - norm_cast at hm, - rw ← finset.coe_Ico at hm, - fin_cases hm; cc, -end - /- `real.log` is strictly concave on $(0, +∞)$. We give an elementary proof rather than using the second derivative test, since this lemma is @@ -372,70 +280,3 @@ begin ... = log (- (a • x + b • y)) : by congr' 1; simp only [algebra.id.smul_eq_mul]; ring ... = _ : by rw log_neg_eq_log, end - -section sqrt_mul_log - -lemma has_deriv_at_sqrt_mul_log {x : ℝ} (hx : x ≠ 0) : - has_deriv_at (λ x, sqrt x * log x) ((2 + log x) / (2 * sqrt x)) x := -begin - convert (has_deriv_at_sqrt hx).mul (has_deriv_at_log hx), - rw [add_div, div_mul_right (sqrt x) two_ne_zero, ←div_eq_mul_inv, sqrt_div_self', - add_comm, div_eq_mul_one_div, mul_comm], -end - -lemma deriv_sqrt_mul_log (x : ℝ) : deriv (λ x, sqrt x * log x) x = (2 + log x) / (2 * sqrt x) := -begin - cases lt_or_le 0 x with hx hx, - { exact (has_deriv_at_sqrt_mul_log hx.ne').deriv }, - { rw [sqrt_eq_zero_of_nonpos hx, mul_zero, div_zero], - refine has_deriv_within_at.deriv_eq_zero _ (unique_diff_on_Iic 0 x hx), - refine (has_deriv_within_at_const x _ 0).congr_of_mem (λ x hx, _) hx, - rw [sqrt_eq_zero_of_nonpos hx, zero_mul] }, -end - -lemma deriv_sqrt_mul_log' : deriv (λ x, sqrt x * log x) = λ x, (2 + log x) / (2 * sqrt x) := -funext deriv_sqrt_mul_log - -lemma deriv2_sqrt_mul_log (x : ℝ) : - deriv^[2] (λ x, sqrt x * log x) x = -log x / (4 * sqrt x ^ 3) := -begin - simp only [nat.iterate, deriv_sqrt_mul_log'], - cases le_or_lt x 0 with hx hx, - { rw [sqrt_eq_zero_of_nonpos hx, zero_pow zero_lt_three, mul_zero, div_zero], - refine has_deriv_within_at.deriv_eq_zero _ (unique_diff_on_Iic 0 x hx), - refine (has_deriv_within_at_const _ _ 0).congr_of_mem (λ x hx, _) hx, - rw [sqrt_eq_zero_of_nonpos hx, mul_zero, div_zero] }, - { have h₀ : sqrt x ≠ 0, from sqrt_ne_zero'.2 hx, - convert (((has_deriv_at_log hx.ne').const_add 2).div - ((has_deriv_at_sqrt hx.ne').const_mul 2) $ mul_ne_zero two_ne_zero h₀).deriv using 1, - nth_rewrite 2 [← mul_self_sqrt hx.le], - field_simp, ring }, -end - -lemma strict_concave_on_sqrt_mul_log_Ioi : strict_concave_on ℝ (set.Ioi 1) (λ x, sqrt x * log x) := -begin - apply strict_concave_on_of_deriv2_neg' (convex_Ioi 1) _ (λ x hx, _), - { exact continuous_sqrt.continuous_on.mul - (continuous_on_log.mono (λ x hx, ne_of_gt (zero_lt_one.trans hx))) }, - { rw [deriv2_sqrt_mul_log x], - exact div_neg_of_neg_of_pos (neg_neg_of_pos (log_pos hx)) - (mul_pos four_pos (pow_pos (sqrt_pos.mpr (zero_lt_one.trans hx)) 3)) }, -end - -end sqrt_mul_log - -open_locale real - -lemma strict_concave_on_sin_Icc : strict_concave_on ℝ (Icc 0 π) sin := -begin - apply strict_concave_on_of_deriv2_neg (convex_Icc _ _) continuous_on_sin (λ x hx, _), - rw interior_Icc at hx, - simp [sin_pos_of_mem_Ioo hx], -end - -lemma strict_concave_on_cos_Icc : strict_concave_on ℝ (Icc (-(π/2)) (π/2)) cos := -begin - apply strict_concave_on_of_deriv2_neg (convex_Icc _ _) continuous_on_cos (λ x hx, _), - rw interior_Icc at hx, - simp [cos_pos_of_mem_Ioo hx], -end diff --git a/src/analysis/convex/specific_functions/deriv.lean b/src/analysis/convex/specific_functions/deriv.lean new file mode 100644 index 0000000000000..c7ba9d71ba84d --- /dev/null +++ b/src/analysis/convex/specific_functions/deriv.lean @@ -0,0 +1,171 @@ +/- +Copyright (c) 2020 Yury Kudryashov. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yury Kudryashov, Sébastien Gouëzel +-/ +import analysis.convex.specific_functions.basic +import analysis.special_functions.pow.deriv +import analysis.special_functions.sqrt +import tactic.linear_combination + +/-! +# Collection of convex functions + +In this file we prove that certain specific functions are strictly convex, including the following: + +* `even.strict_convex_on_pow` : For an even `n : ℕ` with `2 ≤ n`, `λ x, x ^ n` is strictly convex. +* `strict_convex_on_pow` : For `n : ℕ`, with `2 ≤ n`, `λ x, x ^ n` is strictly convex on $[0, +∞)$. +* `strict_convex_on_zpow` : For `m : ℤ` with `m ≠ 0, 1`, `λ x, x ^ m` is strictly convex on + $[0, +∞)$. +* `strict_concave_on_sin_Icc` : `sin` is strictly concave on $[0, π]$ +* `strict_concave_on_cos_Icc` : `cos` is strictly concave on $[-π/2, π/2]$ + +## TODO + +These convexity lemmas are proved by checking the sign of the second derivative. If desired, most +of these could also be switched to elementary proofs, like in +`analysis.convex.specific_functions.basic`. + +-/ + +open real set +open_locale big_operators nnreal + +/-- `x^n`, `n : ℕ` is strictly convex on `[0, +∞)` for all `n` greater than `2`. -/ +lemma strict_convex_on_pow {n : ℕ} (hn : 2 ≤ n) : strict_convex_on ℝ (Ici 0) (λ x : ℝ, x^n) := +begin + apply strict_mono_on.strict_convex_on_of_deriv (convex_Ici _) (continuous_on_pow _), + rw [deriv_pow', interior_Ici], + exact λ x (hx : 0 < x) y hy hxy, mul_lt_mul_of_pos_left (pow_lt_pow_of_lt_left hxy hx.le $ + nat.sub_pos_of_lt hn) (nat.cast_pos.2 $ zero_lt_two.trans_le hn), +end + +/-- `x^n`, `n : ℕ` is strictly convex on the whole real line whenever `n ≠ 0` is even. -/ +lemma even.strict_convex_on_pow {n : ℕ} (hn : even n) (h : n ≠ 0) : + strict_convex_on ℝ set.univ (λ x : ℝ, x^n) := +begin + apply strict_mono.strict_convex_on_univ_of_deriv (continuous_pow n), + rw deriv_pow', + replace h := nat.pos_of_ne_zero h, + exact strict_mono.const_mul (odd.strict_mono_pow $ nat.even.sub_odd h hn $ nat.odd_iff.2 rfl) + (nat.cast_pos.2 h), +end + +lemma finset.prod_nonneg_of_card_nonpos_even + {α β : Type*} [linear_ordered_comm_ring β] + {f : α → β} [decidable_pred (λ x, f x ≤ 0)] + {s : finset α} (h0 : even (s.filter (λ x, f x ≤ 0)).card) : + 0 ≤ ∏ x in s, f x := +calc 0 ≤ (∏ x in s, ((if f x ≤ 0 then (-1:β) else 1) * f x)) : + finset.prod_nonneg (λ x _, by + { split_ifs with hx hx, by simp [hx], simp at hx ⊢, exact le_of_lt hx }) +... = _ : by rw [finset.prod_mul_distrib, finset.prod_ite, finset.prod_const_one, + mul_one, finset.prod_const, neg_one_pow_eq_pow_mod_two, nat.even_iff.1 h0, pow_zero, one_mul] + +lemma int_prod_range_nonneg (m : ℤ) (n : ℕ) (hn : even n) : + 0 ≤ ∏ k in finset.range n, (m - k) := +begin + rcases hn with ⟨n, rfl⟩, + induction n with n ihn, { simp }, + rw ← two_mul at ihn, + rw [← two_mul, nat.succ_eq_add_one, mul_add, mul_one, bit0, ← add_assoc, finset.prod_range_succ, + finset.prod_range_succ, mul_assoc], + refine mul_nonneg ihn _, generalize : (1 + 1) * n = k, + cases le_or_lt m k with hmk hmk, + { have : m ≤ k + 1, from hmk.trans (lt_add_one ↑k).le, + convert mul_nonneg_of_nonpos_of_nonpos (sub_nonpos_of_le hmk) _, + convert sub_nonpos_of_le this }, + { exact mul_nonneg (sub_nonneg_of_le hmk.le) (sub_nonneg_of_le hmk) } +end + +lemma int_prod_range_pos {m : ℤ} {n : ℕ} (hn : even n) (hm : m ∉ Ico (0 : ℤ) n) : + 0 < ∏ k in finset.range n, (m - k) := +begin + refine (int_prod_range_nonneg m n hn).lt_of_ne (λ h, hm _), + rw [eq_comm, finset.prod_eq_zero_iff] at h, + obtain ⟨a, ha, h⟩ := h, + rw sub_eq_zero.1 h, + exact ⟨int.coe_zero_le _, int.coe_nat_lt.2 $ finset.mem_range.1 ha⟩, +end + +/-- `x^m`, `m : ℤ` is convex on `(0, +∞)` for all `m` except `0` and `1`. -/ +lemma strict_convex_on_zpow {m : ℤ} (hm₀ : m ≠ 0) (hm₁ : m ≠ 1) : + strict_convex_on ℝ (Ioi 0) (λ x : ℝ, x^m) := +begin + apply strict_convex_on_of_deriv2_pos' (convex_Ioi 0), + { exact (continuous_on_zpow₀ m).mono (λ x hx, ne_of_gt hx) }, + intros x hx, + rw iter_deriv_zpow, + refine mul_pos _ (zpow_pos_of_pos hx _), + exact_mod_cast int_prod_range_pos (even_bit0 1) (λ hm, _), + norm_cast at hm, + rw ← finset.coe_Ico at hm, + fin_cases hm; cc, +end + +section sqrt_mul_log + +lemma has_deriv_at_sqrt_mul_log {x : ℝ} (hx : x ≠ 0) : + has_deriv_at (λ x, sqrt x * log x) ((2 + log x) / (2 * sqrt x)) x := +begin + convert (has_deriv_at_sqrt hx).mul (has_deriv_at_log hx), + rw [add_div, div_mul_right (sqrt x) two_ne_zero, ←div_eq_mul_inv, sqrt_div_self', + add_comm, div_eq_mul_one_div, mul_comm], +end + +lemma deriv_sqrt_mul_log (x : ℝ) : deriv (λ x, sqrt x * log x) x = (2 + log x) / (2 * sqrt x) := +begin + cases lt_or_le 0 x with hx hx, + { exact (has_deriv_at_sqrt_mul_log hx.ne').deriv }, + { rw [sqrt_eq_zero_of_nonpos hx, mul_zero, div_zero], + refine has_deriv_within_at.deriv_eq_zero _ (unique_diff_on_Iic 0 x hx), + refine (has_deriv_within_at_const x _ 0).congr_of_mem (λ x hx, _) hx, + rw [sqrt_eq_zero_of_nonpos hx, zero_mul] }, +end + +lemma deriv_sqrt_mul_log' : deriv (λ x, sqrt x * log x) = λ x, (2 + log x) / (2 * sqrt x) := +funext deriv_sqrt_mul_log + +lemma deriv2_sqrt_mul_log (x : ℝ) : + deriv^[2] (λ x, sqrt x * log x) x = -log x / (4 * sqrt x ^ 3) := +begin + simp only [nat.iterate, deriv_sqrt_mul_log'], + cases le_or_lt x 0 with hx hx, + { rw [sqrt_eq_zero_of_nonpos hx, zero_pow zero_lt_three, mul_zero, div_zero], + refine has_deriv_within_at.deriv_eq_zero _ (unique_diff_on_Iic 0 x hx), + refine (has_deriv_within_at_const _ _ 0).congr_of_mem (λ x hx, _) hx, + rw [sqrt_eq_zero_of_nonpos hx, mul_zero, div_zero] }, + { have h₀ : sqrt x ≠ 0, from sqrt_ne_zero'.2 hx, + convert (((has_deriv_at_log hx.ne').const_add 2).div + ((has_deriv_at_sqrt hx.ne').const_mul 2) $ mul_ne_zero two_ne_zero h₀).deriv using 1, + nth_rewrite 2 [← mul_self_sqrt hx.le], + field_simp, ring }, +end + +lemma strict_concave_on_sqrt_mul_log_Ioi : strict_concave_on ℝ (set.Ioi 1) (λ x, sqrt x * log x) := +begin + apply strict_concave_on_of_deriv2_neg' (convex_Ioi 1) _ (λ x hx, _), + { exact continuous_sqrt.continuous_on.mul + (continuous_on_log.mono (λ x hx, ne_of_gt (zero_lt_one.trans hx))) }, + { rw [deriv2_sqrt_mul_log x], + exact div_neg_of_neg_of_pos (neg_neg_of_pos (log_pos hx)) + (mul_pos four_pos (pow_pos (sqrt_pos.mpr (zero_lt_one.trans hx)) 3)) }, +end + +end sqrt_mul_log + +open_locale real + +lemma strict_concave_on_sin_Icc : strict_concave_on ℝ (Icc 0 π) sin := +begin + apply strict_concave_on_of_deriv2_neg (convex_Icc _ _) continuous_on_sin (λ x hx, _), + rw interior_Icc at hx, + simp [sin_pos_of_mem_Ioo hx], +end + +lemma strict_concave_on_cos_Icc : strict_concave_on ℝ (Icc (-(π/2)) (π/2)) cos := +begin + apply strict_concave_on_of_deriv2_neg (convex_Icc _ _) continuous_on_cos (λ x hx, _), + rw interior_Icc at hx, + simp [cos_pos_of_mem_Ioo hx], +end diff --git a/src/analysis/fourier/add_circle.lean b/src/analysis/fourier/add_circle.lean index f96ad82907a0d..7b754f1353ec4 100644 --- a/src/analysis/fourier/add_circle.lean +++ b/src/analysis/fourier/add_circle.lean @@ -3,8 +3,8 @@ Copyright (c) 2021 Heather Macbeth. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Heather Macbeth, David Loeffler -/ +import analysis.special_functions.exp_deriv import analysis.special_functions.complex.circle -import topology.instances.add_circle import analysis.inner_product_space.l2_space import measure_theory.function.continuous_map_dense import measure_theory.function.l2_space diff --git a/src/analysis/fourier/fourier_transform.lean b/src/analysis/fourier/fourier_transform.lean index c3dd53ccf025c..dc2b91809dbb5 100644 --- a/src/analysis/fourier/fourier_transform.lean +++ b/src/analysis/fourier/fourier_transform.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: David Loeffler -/ +import analysis.complex.circle import measure_theory.group.integration import measure_theory.measure.haar_of_basis diff --git a/src/analysis/mean_inequalities.lean b/src/analysis/mean_inequalities.lean index af5cfd68579f8..4dba0800e08bd 100644 --- a/src/analysis/mean_inequalities.lean +++ b/src/analysis/mean_inequalities.lean @@ -3,7 +3,9 @@ Copyright (c) 2019 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov, Sébastien Gouëzel, Rémy Degenne -/ -import analysis.convex.specific_functions +import analysis.convex.jensen +import analysis.convex.specific_functions.basic +import analysis.special_functions.pow.nnreal import data.real.conjugate_exponents /-! diff --git a/src/analysis/mean_inequalities_pow.lean b/src/analysis/mean_inequalities_pow.lean index 8292da30dfac0..a694c87faf4f9 100644 --- a/src/analysis/mean_inequalities_pow.lean +++ b/src/analysis/mean_inequalities_pow.lean @@ -3,7 +3,9 @@ Copyright (c) 2019 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov, Sébastien Gouëzel, Rémy Degenne -/ -import analysis.convex.specific_functions +import analysis.convex.jensen +import analysis.convex.specific_functions.basic +import analysis.special_functions.pow.nnreal import tactic.positivity /-! @@ -58,6 +60,23 @@ theorem pow_arith_mean_le_arith_mean_pow_of_even (w z : ι → ℝ) (hw : ∀ i (∑ i in s, w i * z i) ^ n ≤ ∑ i in s, (w i * z i ^ n) := hn.convex_on_pow.map_sum_le hw hw' (λ _ _, trivial) +/-- Specific case of Jensen's inequality for sums of powers -/ +lemma pow_sum_div_card_le_sum_pow {f : ι → ℝ} (n : ℕ) (hf : ∀ a ∈ s, 0 ≤ f a) : + (∑ x in s, f x) ^ (n + 1) / s.card ^ n ≤ ∑ x in s, (f x) ^ (n + 1) := +begin + rcases s.eq_empty_or_nonempty with rfl | hs, + { simp_rw [finset.sum_empty, zero_pow' _ (nat.succ_ne_zero n), zero_div] }, + { have hs0 : 0 < (s.card : ℝ) := nat.cast_pos.2 hs.card_pos, + suffices : (∑ x in s, f x / s.card) ^ (n + 1) ≤ ∑ x in s, (f x ^ (n + 1) / s.card), + { rwa [← finset.sum_div, ← finset.sum_div, div_pow, pow_succ' (s.card : ℝ), + ← div_div, div_le_iff hs0, div_mul, div_self hs0.ne', div_one] at this }, + have := @convex_on.map_sum_le ℝ ℝ ℝ ι _ _ _ _ _ _ (set.Ici 0) (λ x, x ^ (n + 1)) s + (λ _, 1 / s.card) (coe ∘ f) (convex_on_pow (n + 1)) _ _ (λ i hi, set.mem_Ici.2 (hf i hi)), + { simpa only [inv_mul_eq_div, one_div, algebra.id.smul_eq_mul] using this }, + { simp only [one_div, inv_nonneg, nat.cast_nonneg, implies_true_iff] }, + { simpa only [one_div, finset.sum_const, nsmul_eq_mul] using mul_inv_cancel hs0.ne' } } +end + theorem zpow_arith_mean_le_arith_mean_zpow (w z : ι → ℝ) (hw : ∀ i ∈ s, 0 ≤ w i) (hw' : ∑ i in s, w i = 1) (hz : ∀ i ∈ s, 0 < z i) (m : ℤ) : (∑ i in s, w i * z i) ^ m ≤ ∑ i in s, (w i * z i ^ m) := @@ -91,6 +110,11 @@ theorem pow_arith_mean_le_arith_mean_pow (w z : ι → ℝ≥0) (hw' : ∑ i in by exact_mod_cast real.pow_arith_mean_le_arith_mean_pow s _ _ (λ i _, (w i).coe_nonneg) (by exact_mod_cast hw') (λ i _, (z i).coe_nonneg) n +lemma pow_sum_div_card_le_sum_pow (f : ι → ℝ≥0) (n : ℕ) : + (∑ x in s, f x) ^ (n + 1) / s.card ^ n ≤ ∑ x in s, (f x) ^ (n + 1) := +by simpa only [← nnreal.coe_le_coe, nnreal.coe_sum, nonneg.coe_div, nnreal.coe_pow] using + @real.pow_sum_div_card_le_sum_pow ι s (coe ∘ f) n (λ _ _, nnreal.coe_nonneg _) + /-- Weighted generalized mean inequality, version for sums over finite sets, with `ℝ≥0`-valued functions and real exponents. -/ theorem rpow_arith_mean_le_arith_mean_rpow (w z : ι → ℝ≥0) (hw' : ∑ i in s, w i = 1) {p : ℝ} diff --git a/src/analysis/normed_space/lp_space.lean b/src/analysis/normed_space/lp_space.lean index 9cfaeab54c8aa..7ff868c4f74af 100644 --- a/src/analysis/normed_space/lp_space.lean +++ b/src/analysis/normed_space/lp_space.lean @@ -5,6 +5,7 @@ Authors: Heather Macbeth -/ import analysis.mean_inequalities import analysis.mean_inequalities_pow +import analysis.special_functions.pow.continuity import topology.algebra.order.liminf_limsup /-! diff --git a/src/analysis/normed_space/pi_Lp.lean b/src/analysis/normed_space/pi_Lp.lean index ec85d614f1aa0..05a6a28014559 100644 --- a/src/analysis/normed_space/pi_Lp.lean +++ b/src/analysis/normed_space/pi_Lp.lean @@ -5,6 +5,7 @@ Authors: Sébastien Gouëzel, Jireh Loreaux -/ import analysis.mean_inequalities import data.fintype.order +import linear_algebra.matrix.basis /-! # `L^p` distance on finite products of metric spaces diff --git a/src/analysis/special_functions/polar_coord.lean b/src/analysis/special_functions/polar_coord.lean index 05cc1d283346e..18dbbde7ff807 100644 --- a/src/analysis/special_functions/polar_coord.lean +++ b/src/analysis/special_functions/polar_coord.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ +import analysis.special_functions.trigonometric.deriv import measure_theory.function.jacobian /-! diff --git a/src/analysis/special_functions/trigonometric/complex.lean b/src/analysis/special_functions/trigonometric/complex.lean index 5291d37b34aac..35a037e540bd6 100644 --- a/src/analysis/special_functions/trigonometric/complex.lean +++ b/src/analysis/special_functions/trigonometric/complex.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 algebra.quadratic_discriminant -import analysis.convex.specific_functions +import analysis.convex.specific_functions.deriv /-! # Complex trigonometric functions diff --git a/src/geometry/manifold/instances/sphere.lean b/src/geometry/manifold/instances/sphere.lean index 6ab30af5703da..785bef9a7930a 100644 --- a/src/geometry/manifold/instances/sphere.lean +++ b/src/geometry/manifold/instances/sphere.lean @@ -3,8 +3,8 @@ Copyright (c) 2021 Heather Macbeth. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Heather Macbeth -/ -import analysis.complex.circle import analysis.normed_space.ball_action +import analysis.special_functions.exp_deriv import analysis.inner_product_space.calculus import analysis.inner_product_space.pi_L2 import geometry.manifold.algebra.lie_group diff --git a/src/measure_theory/function/jacobian.lean b/src/measure_theory/function/jacobian.lean index 2055a245a40ea..c90328ecd81bd 100644 --- a/src/measure_theory/function/jacobian.lean +++ b/src/measure_theory/function/jacobian.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ +import analysis.calculus.inverse import measure_theory.constructions.borel_space.continuous_linear_map import measure_theory.covering.besicovitch_vector_space import measure_theory.measure.haar_lebesgue diff --git a/src/measure_theory/function/lp_space.lean b/src/measure_theory/function/lp_space.lean index ac8bfcfc86d13..8ce0cf99e6b55 100644 --- a/src/measure_theory/function/lp_space.lean +++ b/src/measure_theory/function/lp_space.lean @@ -5,6 +5,7 @@ Authors: Rémy Degenne, Sébastien Gouëzel -/ import analysis.normed_space.indicator_function import analysis.normed.group.hom +import analysis.special_functions.pow.continuity import measure_theory.function.ess_sup import measure_theory.function.ae_eq_fun import measure_theory.integral.mean_inequalities diff --git a/src/measure_theory/integral/bochner.lean b/src/measure_theory/integral/bochner.lean index 82ee783fd13b7..e2627ba77641b 100644 --- a/src/measure_theory/integral/bochner.lean +++ b/src/measure_theory/integral/bochner.lean @@ -139,6 +139,8 @@ Bochner integral, simple function, function space, Lebesgue dominated convergenc -/ +assert_not_exists differentiable + noncomputable theory open_locale topology big_operators nnreal ennreal measure_theory open set filter topological_space ennreal emetric diff --git a/src/measure_theory/integral/divergence_theorem.lean b/src/measure_theory/integral/divergence_theorem.lean index 17ecec3e87575..465737ffa7277 100644 --- a/src/measure_theory/integral/divergence_theorem.lean +++ b/src/measure_theory/integral/divergence_theorem.lean @@ -5,6 +5,7 @@ Authors: Yury Kudryashov -/ import analysis.box_integral.divergence_theorem import analysis.box_integral.integrability +import analysis.calculus.deriv import measure_theory.integral.interval_integral /-! diff --git a/src/measure_theory/integral/integral_eq_improper.lean b/src/measure_theory/integral/integral_eq_improper.lean index 83d5b5ce1cfdb..b46ea33b9b080 100644 --- a/src/measure_theory/integral/integral_eq_improper.lean +++ b/src/measure_theory/integral/integral_eq_improper.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Anatole Dedecker. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Anatole Dedecker, Bhavik Mehta -/ +import analysis.special_functions.pow.deriv import measure_theory.integral.fund_thm_calculus import order.filter.at_top_bot import measure_theory.function.jacobian diff --git a/src/number_theory/bertrand.lean b/src/number_theory/bertrand.lean index b9dbb2af60d55..a3cde2c5d0267 100644 --- a/src/number_theory/bertrand.lean +++ b/src/number_theory/bertrand.lean @@ -6,7 +6,7 @@ Authors: Patrick Stevens, Bolton Bailey import data.nat.choose.factorization import data.nat.prime_norm_num import number_theory.primorial -import analysis.convex.specific_functions +import analysis.convex.specific_functions.deriv /-! # Bertrand's Postulate diff --git a/src/topology/metric_space/hausdorff_dimension.lean b/src/topology/metric_space/hausdorff_dimension.lean index da666adf355ce..2dee87d9172dd 100644 --- a/src/topology/metric_space/hausdorff_dimension.lean +++ b/src/topology/metric_space/hausdorff_dimension.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ +import analysis.calculus.cont_diff import measure_theory.measure.hausdorff /-! From 5a4ea8453f128345f73cc656e80a49de2a54f481 Mon Sep 17 00:00:00 2001 From: Chris Hughes Date: Thu, 18 May 2023 15:19:24 +0000 Subject: [PATCH 1032/1291] chore(data/set/Union_lift) generalize Union_lift to Sort (#19033) --- src/data/set/Union_lift.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/data/set/Union_lift.lean b/src/data/set/Union_lift.lean index 80b88373cbbed..017854b0d7b2d 100644 --- a/src/data/set/Union_lift.lean +++ b/src/data/set/Union_lift.lean @@ -38,7 +38,7 @@ constants, unary functions, or binary functions are preserved. These lemmas are: directed union, directed supremum, glue, gluing -/ -variables {α ι β : Type*} +variables {α : Type*} {ι β : Sort*} namespace set From 74403a3b2551b0970855e14ef5e8fd0d6af1bfc2 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Thu, 18 May 2023 17:57:58 +0000 Subject: [PATCH 1033/1291] chore(algebra/category/Module/monoidal): split file (#19034) Co-authored-by: Scott Morrison --- src/algebra/category/Module/adjunctions.lean | 2 +- .../{monoidal.lean => monoidal/basic.lean} | 125 +----------------- .../category/Module/monoidal/closed.lean | 87 ++++++++++++ .../category/Module/monoidal/symmetric.lean | 72 ++++++++++ src/algebra/category/fgModule/basic.lean | 2 +- .../monoidal/internal/Module.lean | 2 +- src/representation_theory/Rep.lean | 2 +- 7 files changed, 169 insertions(+), 123 deletions(-) rename src/algebra/category/Module/{monoidal.lean => monoidal/basic.lean} (64%) create mode 100644 src/algebra/category/Module/monoidal/closed.lean create mode 100644 src/algebra/category/Module/monoidal/symmetric.lean diff --git a/src/algebra/category/Module/adjunctions.lean b/src/algebra/category/Module/adjunctions.lean index e03b0271ee9ce..45e40c8bccec8 100644 --- a/src/algebra/category/Module/adjunctions.lean +++ b/src/algebra/category/Module/adjunctions.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison, Johan Commelin -/ -import algebra.category.Module.monoidal +import algebra.category.Module.monoidal.basic import category_theory.monoidal.functorial import category_theory.monoidal.types import linear_algebra.direct_sum.finsupp diff --git a/src/algebra/category/Module/monoidal.lean b/src/algebra/category/Module/monoidal/basic.lean similarity index 64% rename from src/algebra/category/Module/monoidal.lean rename to src/algebra/category/Module/monoidal/basic.lean index 20dd4ea2f8fc4..670fb7c145ad0 100644 --- a/src/algebra/category/Module/monoidal.lean +++ b/src/algebra/category/Module/monoidal/basic.lean @@ -3,25 +3,26 @@ Copyright (c) 2020 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kevin Buzzard, Scott Morrison, Jakob von Raumer -/ -import category_theory.monoidal.braided -import category_theory.closed.monoidal import algebra.category.Module.basic import linear_algebra.tensor_product import category_theory.linear.yoneda import category_theory.monoidal.linear /-! -# The symmetric monoidal category structure on R-modules +# The monoidal category structure on R-modules Mostly this uses existing machinery in `linear_algebra.tensor_product`. We just need to provide a few small missing pieces to build the -`monoidal_category` instance and then the `symmetric_category` instance. +`monoidal_category` instance. +The `symmetric_category` instance is in `algebra.category.Module.monoidal.symmetric` +to reduce imports. Note the universe level of the modules must be at least the universe level of the ring, so that we have a monoidal unit. For now, we simplify by insisting both universe levels are the same. -We then construct the monoidal closed structure on `Module R`. +We construct the monoidal closed structure on `Module R` in +`algebra.category.Module.monoidal.closed`. If you're happy using the bundled `Module R`, it may be possible to mostly use this as an interface and not need to interact much with the implementation details. @@ -215,57 +216,6 @@ lemma associator_inv_apply {M N K : Module.{u} R} (m : M) (n : N) (k : K) : end monoidal_category -/-- (implementation) the braiding for R-modules -/ -def braiding (M N : Module R) : tensor_obj M N ≅ tensor_obj N M := -linear_equiv.to_Module_iso (tensor_product.comm R M N) - -@[simp] lemma braiding_naturality {X₁ X₂ Y₁ Y₂ : Module.{u} R} (f : X₁ ⟶ Y₁) (g : X₂ ⟶ Y₂) : - (f ⊗ g) ≫ (Y₁.braiding Y₂).hom = - (X₁.braiding X₂).hom ≫ (g ⊗ f) := -begin - apply tensor_product.ext', - intros x y, - refl -end - -@[simp] lemma hexagon_forward (X Y Z : Module.{u} R) : - (α_ X Y Z).hom ≫ (braiding X _).hom ≫ (α_ Y Z X).hom = - ((braiding X Y).hom ⊗ 𝟙 Z) ≫ (α_ Y X Z).hom ≫ (𝟙 Y ⊗ (braiding X Z).hom) := -begin - apply tensor_product.ext_threefold, - intros x y z, - refl, -end - -@[simp] lemma hexagon_reverse (X Y Z : Module.{u} R) : - (α_ X Y Z).inv ≫ (braiding _ Z).hom ≫ (α_ Z X Y).inv = - (𝟙 X ⊗ (Y.braiding Z).hom) ≫ (α_ X Z Y).inv ≫ ((X.braiding Z).hom ⊗ 𝟙 Y) := -begin - apply (cancel_epi (α_ X Y Z).hom).1, - apply tensor_product.ext_threefold, - intros x y z, - refl, -end - -local attribute [ext] tensor_product.ext - -/-- The symmetric monoidal structure on `Module R`. -/ -instance symmetric_category : symmetric_category (Module.{u} R) := -{ braiding := braiding, - braiding_naturality' := λ X₁ X₂ Y₁ Y₂ f g, braiding_naturality f g, - hexagon_forward' := hexagon_forward, - hexagon_reverse' := hexagon_reverse, } - -namespace monoidal_category - -@[simp] lemma braiding_hom_apply {M N : Module.{u} R} (m : M) (n : N) : - ((β_ M N).hom : M ⊗ N ⟶ N ⊗ M) (m ⊗ₜ n) = n ⊗ₜ m := rfl - -@[simp] lemma braiding_inv_apply {M N : Module.{u} R} (m : M) (n : N) : - ((β_ M N).inv : N ⊗ M ⟶ M ⊗ N) (n ⊗ₜ m) = m ⊗ₜ n := rfl - -end monoidal_category - open opposite instance : monoidal_preadditive (Module.{u} R) := @@ -281,67 +231,4 @@ by refine ⟨_, _⟩; dsimp only [auto_param]; intros; simp only [linear_map.compr₂_apply, tensor_product.mk_apply, monoidal_category.hom_apply, linear_map.smul_apply, tensor_product.tmul_smul, tensor_product.smul_tmul] -/-- -Auxiliary definition for the `monoidal_closed` instance on `Module R`. -(This is only a separate definition in order to speed up typechecking. ) --/ -@[simps] -def monoidal_closed_hom_equiv (M N P : Module.{u} R) : - ((monoidal_category.tensor_left M).obj N ⟶ P) ≃ - (N ⟶ ((linear_coyoneda R (Module R)).obj (op M)).obj P) := -{ to_fun := λ f, linear_map.compr₂ (tensor_product.mk R N M) ((β_ N M).hom ≫ f), - inv_fun := λ f, (β_ M N).hom ≫ tensor_product.lift f, - left_inv := λ f, begin ext m n, - simp only [tensor_product.mk_apply, tensor_product.lift.tmul, linear_map.compr₂_apply, - function.comp_app, coe_comp, monoidal_category.braiding_hom_apply], - end, - right_inv := λ f, begin ext m n, - simp only [tensor_product.mk_apply, tensor_product.lift.tmul, linear_map.compr₂_apply, - symmetric_category.symmetry_assoc], - end, } - -instance : monoidal_closed (Module.{u} R) := -{ closed' := λ M, - { is_adj := - { right := (linear_coyoneda R (Module.{u} R)).obj (op M), - adj := adjunction.mk_of_hom_equiv - { hom_equiv := λ N P, monoidal_closed_hom_equiv M N P, } } } } - -lemma ihom_map_apply {M N P : Module.{u} R} (f : N ⟶ P) (g : Module.of R (M ⟶ N)) : - (ihom M).map f g = g ≫ f := rfl - --- I can't seem to express the function coercion here without writing `@coe_fn`. -@[simp] -lemma monoidal_closed_curry {M N P : Module.{u} R} (f : M ⊗ N ⟶ P) (x : M) (y : N) : - @coe_fn _ _ linear_map.has_coe_to_fun ((monoidal_closed.curry f : N →ₗ[R] (M →ₗ[R] P)) y) x = - f (x ⊗ₜ[R] y) := -rfl - -@[simp] -lemma monoidal_closed_uncurry {M N P : Module.{u} R} - (f : N ⟶ (M ⟶[Module.{u} R] P)) (x : M) (y : N) : - monoidal_closed.uncurry f (x ⊗ₜ[R] y) = (@coe_fn _ _ linear_map.has_coe_to_fun (f y)) x := -rfl - -/-- Describes the counit of the adjunction `M ⊗ - ⊣ Hom(M, -)`. Given an `R`-module `N` this -should give a map `M ⊗ Hom(M, N) ⟶ N`, so we flip the order of the arguments in the identity map -`Hom(M, N) ⟶ (M ⟶ N)` and uncurry the resulting map `M ⟶ Hom(M, N) ⟶ N.` -/ -lemma ihom_ev_app (M N : Module.{u} R) : - (ihom.ev M).app N = tensor_product.uncurry _ _ _ _ linear_map.id.flip := -begin - ext, - exact Module.monoidal_closed_uncurry _ _ _, -end - -/-- Describes the unit of the adjunction `M ⊗ - ⊣ Hom(M, -)`. Given an `R`-module `N` this should -define a map `N ⟶ Hom(M, M ⊗ N)`, which is given by flipping the arguments in the natural -`R`-bilinear map `M ⟶ N ⟶ M ⊗ N`. -/ -lemma ihom_coev_app (M N : Module.{u} R) : - (ihom.coev M).app N = (tensor_product.mk _ _ _).flip := -rfl - -lemma monoidal_closed_pre_app {M N : Module.{u} R} (P : Module.{u} R) (f : N ⟶ M) : - (monoidal_closed.pre f).app P = linear_map.lcomp R _ f := -rfl - end Module diff --git a/src/algebra/category/Module/monoidal/closed.lean b/src/algebra/category/Module/monoidal/closed.lean new file mode 100644 index 0000000000000..498cff23a9b50 --- /dev/null +++ b/src/algebra/category/Module/monoidal/closed.lean @@ -0,0 +1,87 @@ +/- +Copyright (c) 2020 Scott Morrison. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Kevin Buzzard, Scott Morrison, Jakob von Raumer +-/ +import category_theory.closed.monoidal +import algebra.category.Module.monoidal.symmetric + +/-! +# The monoidal closed structure on `Module R`. +-/ + +universes v w x u + +open category_theory +open opposite + +namespace Module + +variables {R : Type u} [comm_ring R] + +local attribute [ext] tensor_product.ext + +/-- +Auxiliary definition for the `monoidal_closed` instance on `Module R`. +(This is only a separate definition in order to speed up typechecking. ) +-/ +@[simps] +def monoidal_closed_hom_equiv (M N P : Module.{u} R) : + ((monoidal_category.tensor_left M).obj N ⟶ P) ≃ + (N ⟶ ((linear_coyoneda R (Module R)).obj (op M)).obj P) := +{ to_fun := λ f, linear_map.compr₂ (tensor_product.mk R N M) ((β_ N M).hom ≫ f), + inv_fun := λ f, (β_ M N).hom ≫ tensor_product.lift f, + left_inv := λ f, begin ext m n, + simp only [tensor_product.mk_apply, tensor_product.lift.tmul, linear_map.compr₂_apply, + function.comp_app, coe_comp, monoidal_category.braiding_hom_apply], + end, + right_inv := λ f, begin ext m n, + simp only [tensor_product.mk_apply, tensor_product.lift.tmul, linear_map.compr₂_apply, + symmetric_category.symmetry_assoc], + end, } + +instance : monoidal_closed (Module.{u} R) := +{ closed' := λ M, + { is_adj := + { right := (linear_coyoneda R (Module.{u} R)).obj (op M), + adj := adjunction.mk_of_hom_equiv + { hom_equiv := λ N P, monoidal_closed_hom_equiv M N P, } } } } + +lemma ihom_map_apply {M N P : Module.{u} R} (f : N ⟶ P) (g : Module.of R (M ⟶ N)) : + (ihom M).map f g = g ≫ f := rfl + +-- I can't seem to express the function coercion here without writing `@coe_fn`. +@[simp] +lemma monoidal_closed_curry {M N P : Module.{u} R} (f : M ⊗ N ⟶ P) (x : M) (y : N) : + @coe_fn _ _ linear_map.has_coe_to_fun ((monoidal_closed.curry f : N →ₗ[R] (M →ₗ[R] P)) y) x = + f (x ⊗ₜ[R] y) := +rfl + +@[simp] +lemma monoidal_closed_uncurry {M N P : Module.{u} R} + (f : N ⟶ (M ⟶[Module.{u} R] P)) (x : M) (y : N) : + monoidal_closed.uncurry f (x ⊗ₜ[R] y) = (@coe_fn _ _ linear_map.has_coe_to_fun (f y)) x := +rfl + +/-- Describes the counit of the adjunction `M ⊗ - ⊣ Hom(M, -)`. Given an `R`-module `N` this +should give a map `M ⊗ Hom(M, N) ⟶ N`, so we flip the order of the arguments in the identity map +`Hom(M, N) ⟶ (M ⟶ N)` and uncurry the resulting map `M ⟶ Hom(M, N) ⟶ N.` -/ +lemma ihom_ev_app (M N : Module.{u} R) : + (ihom.ev M).app N = tensor_product.uncurry _ _ _ _ linear_map.id.flip := +begin + ext, + exact Module.monoidal_closed_uncurry _ _ _, +end + +/-- Describes the unit of the adjunction `M ⊗ - ⊣ Hom(M, -)`. Given an `R`-module `N` this should +define a map `N ⟶ Hom(M, M ⊗ N)`, which is given by flipping the arguments in the natural +`R`-bilinear map `M ⟶ N ⟶ M ⊗ N`. -/ +lemma ihom_coev_app (M N : Module.{u} R) : + (ihom.coev M).app N = (tensor_product.mk _ _ _).flip := +rfl + +lemma monoidal_closed_pre_app {M N : Module.{u} R} (P : Module.{u} R) (f : N ⟶ M) : + (monoidal_closed.pre f).app P = linear_map.lcomp R _ f := +rfl + +end Module diff --git a/src/algebra/category/Module/monoidal/symmetric.lean b/src/algebra/category/Module/monoidal/symmetric.lean new file mode 100644 index 0000000000000..9bf57b263f96c --- /dev/null +++ b/src/algebra/category/Module/monoidal/symmetric.lean @@ -0,0 +1,72 @@ +/- +Copyright (c) 2020 Scott Morrison. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Kevin Buzzard, Scott Morrison, Jakob von Raumer +-/ +import category_theory.monoidal.braided +import algebra.category.Module.monoidal.basic + +/-! +# The symmetric monoidal structure on `Module R`. +-/ + +universes v w x u + +open category_theory + +namespace Module + +variables {R : Type u} [comm_ring R] + +/-- (implementation) the braiding for R-modules -/ +def braiding (M N : Module.{u} R) : (M ⊗ N) ≅ (N ⊗ M) := +linear_equiv.to_Module_iso (tensor_product.comm R M N) + +namespace monoidal_category + +@[simp] lemma braiding_naturality {X₁ X₂ Y₁ Y₂ : Module.{u} R} (f : X₁ ⟶ Y₁) (g : X₂ ⟶ Y₂) : + (f ⊗ g) ≫ (Y₁.braiding Y₂).hom = + (X₁.braiding X₂).hom ≫ (g ⊗ f) := +begin + apply tensor_product.ext', + intros x y, + refl +end + +@[simp] lemma hexagon_forward (X Y Z : Module.{u} R) : + (α_ X Y Z).hom ≫ (braiding X _).hom ≫ (α_ Y Z X).hom = + ((braiding X Y).hom ⊗ 𝟙 Z) ≫ (α_ Y X Z).hom ≫ (𝟙 Y ⊗ (braiding X Z).hom) := +begin + apply tensor_product.ext_threefold, + intros x y z, + refl, +end + +@[simp] lemma hexagon_reverse (X Y Z : Module.{u} R) : + (α_ X Y Z).inv ≫ (braiding _ Z).hom ≫ (α_ Z X Y).inv = + (𝟙 X ⊗ (Y.braiding Z).hom) ≫ (α_ X Z Y).inv ≫ ((X.braiding Z).hom ⊗ 𝟙 Y) := +begin + apply (cancel_epi (α_ X Y Z).hom).1, + apply tensor_product.ext_threefold, + intros x y z, + refl, +end + +local attribute [ext] tensor_product.ext + +/-- The symmetric monoidal structure on `Module R`. -/ +instance symmetric_category : symmetric_category (Module.{u} R) := +{ braiding := braiding, + braiding_naturality' := λ X₁ X₂ Y₁ Y₂ f g, braiding_naturality f g, + hexagon_forward' := hexagon_forward, + hexagon_reverse' := hexagon_reverse, } + +@[simp] lemma braiding_hom_apply {M N : Module.{u} R} (m : M) (n : N) : + ((β_ M N).hom : M ⊗ N ⟶ N ⊗ M) (m ⊗ₜ n) = n ⊗ₜ m := rfl + +@[simp] lemma braiding_inv_apply {M N : Module.{u} R} (m : M) (n : N) : + ((β_ M N).inv : N ⊗ M ⟶ M ⊗ N) (n ⊗ₜ m) = m ⊗ₜ n := rfl + +end monoidal_category + +end Module diff --git a/src/algebra/category/fgModule/basic.lean b/src/algebra/category/fgModule/basic.lean index c439cce1df2ee..f70d4ec8e3ecb 100644 --- a/src/algebra/category/fgModule/basic.lean +++ b/src/algebra/category/fgModule/basic.lean @@ -7,7 +7,7 @@ import category_theory.monoidal.rigid.basic import category_theory.monoidal.subcategory import linear_algebra.coevaluation import linear_algebra.free_module.finite.matrix -import algebra.category.Module.monoidal +import algebra.category.Module.monoidal.closed /-! # The category of finitely generated modules over a ring diff --git a/src/category_theory/monoidal/internal/Module.lean b/src/category_theory/monoidal/internal/Module.lean index ba73dec00fd2a..27071854b7d83 100644 --- a/src/category_theory/monoidal/internal/Module.lean +++ b/src/category_theory/monoidal/internal/Module.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ -import algebra.category.Module.monoidal +import algebra.category.Module.monoidal.basic import algebra.category.Algebra.basic import category_theory.monoidal.Mon_ diff --git a/src/representation_theory/Rep.lean b/src/representation_theory/Rep.lean index b0cf5eea56fbe..01927f4e3b3dd 100644 --- a/src/representation_theory/Rep.lean +++ b/src/representation_theory/Rep.lean @@ -7,7 +7,7 @@ import representation_theory.basic import representation_theory.Action import algebra.category.Module.abelian import algebra.category.Module.colimits -import algebra.category.Module.monoidal +import algebra.category.Module.monoidal.closed import algebra.category.Module.adjunctions import category_theory.closed.functor_category From 95a87616d63b3cb49d3fe678d416fbe9c4217bf4 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Thu, 18 May 2023 17:57:59 +0000 Subject: [PATCH 1034/1291] chore(category_theory/monoidal/types): split file (#19035) Co-authored-by: Scott Morrison --- src/algebra/category/Module/adjunctions.lean | 2 +- src/category_theory/enriched/basic.lean | 3 +- .../monoidal/internal/types.lean | 2 +- .../basic.lean} | 72 +----------- .../of_chosen_finite_products/symmetric.lean | 111 ++++++++++++++++++ .../monoidal/{types.lean => types/basic.lean} | 41 +------ .../monoidal/types/coyoneda.lean | 45 +++++++ .../monoidal/types/symmetric.lean | 28 +++++ src/representation_theory/Action.lean | 2 +- 9 files changed, 194 insertions(+), 112 deletions(-) rename src/category_theory/monoidal/{of_chosen_finite_products.lean => of_chosen_finite_products/basic.lean} (79%) create mode 100644 src/category_theory/monoidal/of_chosen_finite_products/symmetric.lean rename src/category_theory/monoidal/{types.lean => types/basic.lean} (58%) create mode 100644 src/category_theory/monoidal/types/coyoneda.lean create mode 100644 src/category_theory/monoidal/types/symmetric.lean diff --git a/src/algebra/category/Module/adjunctions.lean b/src/algebra/category/Module/adjunctions.lean index 45e40c8bccec8..b2d2ceb2793cf 100644 --- a/src/algebra/category/Module/adjunctions.lean +++ b/src/algebra/category/Module/adjunctions.lean @@ -5,7 +5,7 @@ Authors: Scott Morrison, Johan Commelin -/ import algebra.category.Module.monoidal.basic import category_theory.monoidal.functorial -import category_theory.monoidal.types +import category_theory.monoidal.types.basic import linear_algebra.direct_sum.finsupp import category_theory.linear.linear_functor diff --git a/src/category_theory/enriched/basic.lean b/src/category_theory/enriched/basic.lean index d7cd99921aebb..1510eb07d5e5d 100644 --- a/src/category_theory/enriched/basic.lean +++ b/src/category_theory/enriched/basic.lean @@ -3,7 +3,8 @@ Copyright (c) 2021 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ -import category_theory.monoidal.types +import category_theory.monoidal.types.symmetric +import category_theory.monoidal.types.coyoneda import category_theory.monoidal.center import tactic.apply_fun diff --git a/src/category_theory/monoidal/internal/types.lean b/src/category_theory/monoidal/internal/types.lean index 0830caa718827..6351e83a93dac 100644 --- a/src/category_theory/monoidal/internal/types.lean +++ b/src/category_theory/monoidal/internal/types.lean @@ -5,7 +5,7 @@ Authors: Scott Morrison -/ import algebra.category.Mon.basic import category_theory.monoidal.CommMon_ -import category_theory.monoidal.types +import category_theory.monoidal.types.symmetric /-! # `Mon_ (Type u) ≌ Mon.{u}` diff --git a/src/category_theory/monoidal/of_chosen_finite_products.lean b/src/category_theory/monoidal/of_chosen_finite_products/basic.lean similarity index 79% rename from src/category_theory/monoidal/of_chosen_finite_products.lean rename to src/category_theory/monoidal/of_chosen_finite_products/basic.lean index e7cbcbc84114f..93e166e639cc6 100644 --- a/src/category_theory/monoidal/of_chosen_finite_products.lean +++ b/src/category_theory/monoidal/of_chosen_finite_products/basic.lean @@ -3,7 +3,7 @@ Copyright (c) 2019 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison, Simon Hudon -/ -import category_theory.monoidal.braided +import category_theory.monoidal.category import category_theory.limits.shapes.binary_products import category_theory.pempty @@ -327,78 +327,8 @@ def monoidal_of_chosen_finite_products_synonym instance : monoidal_category (monoidal_of_chosen_finite_products_synonym 𝒯 ℬ) := monoidal_of_chosen_finite_products 𝒯 ℬ -lemma braiding_naturality {X X' Y Y' : C} (f : X ⟶ Y) (g : X' ⟶ Y') : - (tensor_hom ℬ f g) ≫ (limits.binary_fan.braiding (ℬ Y Y').is_limit (ℬ Y' Y).is_limit).hom = - (limits.binary_fan.braiding (ℬ X X').is_limit (ℬ X' X).is_limit).hom ≫ (tensor_hom ℬ g f) := -begin - dsimp [tensor_hom, limits.binary_fan.braiding], - apply (ℬ _ _).is_limit.hom_ext, rintro ⟨⟨⟩⟩; - { dsimp [limits.is_limit.cone_point_unique_up_to_iso], simp, }, -end - -lemma hexagon_forward (X Y Z : C) : - (binary_fan.associator_of_limit_cone ℬ X Y Z).hom ≫ - (limits.binary_fan.braiding - (ℬ X (tensor_obj ℬ Y Z)).is_limit - (ℬ (tensor_obj ℬ Y Z) X).is_limit).hom ≫ - (binary_fan.associator_of_limit_cone ℬ Y Z X).hom = - (tensor_hom ℬ (limits.binary_fan.braiding (ℬ X Y).is_limit (ℬ Y X).is_limit).hom (𝟙 Z)) ≫ - (binary_fan.associator_of_limit_cone ℬ Y X Z).hom ≫ - (tensor_hom ℬ (𝟙 Y) (limits.binary_fan.braiding (ℬ X Z).is_limit (ℬ Z X).is_limit).hom) := -begin - dsimp [tensor_hom, limits.binary_fan.braiding], - apply (ℬ _ _).is_limit.hom_ext, rintro ⟨⟨⟩⟩, - { dsimp [limits.is_limit.cone_point_unique_up_to_iso], simp, }, - { apply (ℬ _ _).is_limit.hom_ext, rintro ⟨⟨⟩⟩; - { dsimp [limits.is_limit.cone_point_unique_up_to_iso], simp, }, } -end - -lemma hexagon_reverse (X Y Z : C) : - (binary_fan.associator_of_limit_cone ℬ X Y Z).inv ≫ - (limits.binary_fan.braiding - (ℬ (tensor_obj ℬ X Y) Z).is_limit - (ℬ Z (tensor_obj ℬ X Y)).is_limit).hom ≫ - (binary_fan.associator_of_limit_cone ℬ Z X Y).inv = - (tensor_hom ℬ (𝟙 X) (limits.binary_fan.braiding (ℬ Y Z).is_limit (ℬ Z Y).is_limit).hom) ≫ - (binary_fan.associator_of_limit_cone ℬ X Z Y).inv ≫ - (tensor_hom ℬ (limits.binary_fan.braiding (ℬ X Z).is_limit (ℬ Z X).is_limit).hom (𝟙 Y)) := -begin - dsimp [tensor_hom, limits.binary_fan.braiding], - apply (ℬ _ _).is_limit.hom_ext, rintro ⟨⟨⟩⟩, - { apply (ℬ _ _).is_limit.hom_ext, rintro ⟨⟨⟩⟩; - { dsimp [binary_fan.associator_of_limit_cone, binary_fan.associator, - limits.is_limit.cone_point_unique_up_to_iso], - simp, }, }, - { dsimp [binary_fan.associator_of_limit_cone, binary_fan.associator, - limits.is_limit.cone_point_unique_up_to_iso], - simp, }, -end - -lemma symmetry (X Y : C) : - (limits.binary_fan.braiding (ℬ X Y).is_limit (ℬ Y X).is_limit).hom ≫ - (limits.binary_fan.braiding (ℬ Y X).is_limit (ℬ X Y).is_limit).hom = - 𝟙 (tensor_obj ℬ X Y) := -begin - dsimp [tensor_hom, limits.binary_fan.braiding], - apply (ℬ _ _).is_limit.hom_ext, rintro ⟨⟨⟩⟩; - { dsimp [limits.is_limit.cone_point_unique_up_to_iso], simp, }, -end - end monoidal_of_chosen_finite_products -open monoidal_of_chosen_finite_products - -/-- -The monoidal structure coming from finite products is symmetric. --/ -def symmetric_of_chosen_finite_products : - symmetric_category (monoidal_of_chosen_finite_products_synonym 𝒯 ℬ) := -{ braiding := λ X Y, limits.binary_fan.braiding (ℬ _ _).is_limit (ℬ _ _).is_limit, - braiding_naturality' := λ X X' Y Y' f g, braiding_naturality ℬ f g, - hexagon_forward' := λ X Y Z, hexagon_forward ℬ X Y Z, - hexagon_reverse' := λ X Y Z, hexagon_reverse ℬ X Y Z, - symmetry' := λ X Y, symmetry ℬ X Y, } - end end category_theory diff --git a/src/category_theory/monoidal/of_chosen_finite_products/symmetric.lean b/src/category_theory/monoidal/of_chosen_finite_products/symmetric.lean new file mode 100644 index 0000000000000..7200c53a49fd8 --- /dev/null +++ b/src/category_theory/monoidal/of_chosen_finite_products/symmetric.lean @@ -0,0 +1,111 @@ +/- +Copyright (c) 2019 Scott Morrison. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Scott Morrison, Simon Hudon +-/ +import category_theory.monoidal.braided +import category_theory.monoidal.of_chosen_finite_products.basic + +/-! +# The symmetric monoidal structure on a category with chosen finite products. + +-/ + +universes v u + +noncomputable theory + +namespace category_theory + +variables (C : Type u) [category.{v} C] {X Y : C} + +open category_theory.limits + +section +local attribute [tidy] tactic.case_bash + +variables {C} +variables (𝒯 : limit_cone (functor.empty.{v} C)) +variables (ℬ : Π (X Y : C), limit_cone (pair X Y)) + +open monoidal_of_chosen_finite_products + +namespace monoidal_of_chosen_finite_products + +open monoidal_category + +lemma braiding_naturality {X X' Y Y' : C} (f : X ⟶ Y) (g : X' ⟶ Y') : + (tensor_hom ℬ f g) ≫ (limits.binary_fan.braiding (ℬ Y Y').is_limit (ℬ Y' Y).is_limit).hom = + (limits.binary_fan.braiding (ℬ X X').is_limit (ℬ X' X).is_limit).hom ≫ (tensor_hom ℬ g f) := +begin + dsimp [tensor_hom, limits.binary_fan.braiding], + apply (ℬ _ _).is_limit.hom_ext, rintro ⟨⟨⟩⟩; + { dsimp [limits.is_limit.cone_point_unique_up_to_iso], simp, }, +end + +lemma hexagon_forward (X Y Z : C) : + (binary_fan.associator_of_limit_cone ℬ X Y Z).hom ≫ + (limits.binary_fan.braiding + (ℬ X (tensor_obj ℬ Y Z)).is_limit + (ℬ (tensor_obj ℬ Y Z) X).is_limit).hom ≫ + (binary_fan.associator_of_limit_cone ℬ Y Z X).hom = + (tensor_hom ℬ (limits.binary_fan.braiding (ℬ X Y).is_limit (ℬ Y X).is_limit).hom (𝟙 Z)) ≫ + (binary_fan.associator_of_limit_cone ℬ Y X Z).hom ≫ + (tensor_hom ℬ (𝟙 Y) (limits.binary_fan.braiding (ℬ X Z).is_limit (ℬ Z X).is_limit).hom) := +begin + dsimp [tensor_hom, limits.binary_fan.braiding], + apply (ℬ _ _).is_limit.hom_ext, rintro ⟨⟨⟩⟩, + { dsimp [limits.is_limit.cone_point_unique_up_to_iso], simp, }, + { apply (ℬ _ _).is_limit.hom_ext, rintro ⟨⟨⟩⟩; + { dsimp [limits.is_limit.cone_point_unique_up_to_iso], simp, }, } +end + +lemma hexagon_reverse (X Y Z : C) : + (binary_fan.associator_of_limit_cone ℬ X Y Z).inv ≫ + (limits.binary_fan.braiding + (ℬ (tensor_obj ℬ X Y) Z).is_limit + (ℬ Z (tensor_obj ℬ X Y)).is_limit).hom ≫ + (binary_fan.associator_of_limit_cone ℬ Z X Y).inv = + (tensor_hom ℬ (𝟙 X) (limits.binary_fan.braiding (ℬ Y Z).is_limit (ℬ Z Y).is_limit).hom) ≫ + (binary_fan.associator_of_limit_cone ℬ X Z Y).inv ≫ + (tensor_hom ℬ (limits.binary_fan.braiding (ℬ X Z).is_limit (ℬ Z X).is_limit).hom (𝟙 Y)) := +begin + dsimp [tensor_hom, limits.binary_fan.braiding], + apply (ℬ _ _).is_limit.hom_ext, rintro ⟨⟨⟩⟩, + { apply (ℬ _ _).is_limit.hom_ext, rintro ⟨⟨⟩⟩; + { dsimp [binary_fan.associator_of_limit_cone, binary_fan.associator, + limits.is_limit.cone_point_unique_up_to_iso], + simp, }, }, + { dsimp [binary_fan.associator_of_limit_cone, binary_fan.associator, + limits.is_limit.cone_point_unique_up_to_iso], + simp, }, +end + +lemma symmetry (X Y : C) : + (limits.binary_fan.braiding (ℬ X Y).is_limit (ℬ Y X).is_limit).hom ≫ + (limits.binary_fan.braiding (ℬ Y X).is_limit (ℬ X Y).is_limit).hom = + 𝟙 (tensor_obj ℬ X Y) := +begin + dsimp [tensor_hom, limits.binary_fan.braiding], + apply (ℬ _ _).is_limit.hom_ext, rintro ⟨⟨⟩⟩; + { dsimp [limits.is_limit.cone_point_unique_up_to_iso], simp, }, +end + +end monoidal_of_chosen_finite_products + +open monoidal_of_chosen_finite_products + +/-- +The monoidal structure coming from finite products is symmetric. +-/ +def symmetric_of_chosen_finite_products : + symmetric_category (monoidal_of_chosen_finite_products_synonym 𝒯 ℬ) := +{ braiding := λ X Y, limits.binary_fan.braiding (ℬ _ _).is_limit (ℬ _ _).is_limit, + braiding_naturality' := λ X X' Y Y' f g, braiding_naturality ℬ f g, + hexagon_forward' := λ X Y Z, hexagon_forward ℬ X Y Z, + hexagon_reverse' := λ X Y Z, hexagon_reverse ℬ X Y Z, + symmetry' := λ X Y, symmetry ℬ X Y, } + +end + +end category_theory diff --git a/src/category_theory/monoidal/types.lean b/src/category_theory/monoidal/types/basic.lean similarity index 58% rename from src/category_theory/monoidal/types.lean rename to src/category_theory/monoidal/types/basic.lean index a96a3d16be555..84bcfd68e41f7 100644 --- a/src/category_theory/monoidal/types.lean +++ b/src/category_theory/monoidal/types/basic.lean @@ -3,12 +3,13 @@ Copyright (c) 2018 Michael Jendrusch. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Michael Jendrusch, Scott Morrison -/ -import category_theory.monoidal.of_chosen_finite_products +import category_theory.monoidal.functor +import category_theory.monoidal.of_chosen_finite_products.basic import category_theory.limits.shapes.types import logic.equiv.fin /-! -# The category of types is a symmetric monoidal category +# The category of types is a monoidal category -/ open category_theory @@ -22,9 +23,6 @@ namespace category_theory instance types_monoidal : monoidal_category.{u} (Type u) := monoidal_of_chosen_finite_products (types.terminal_limit_cone) (types.binary_product_limit_cone) -instance types_symmetric : symmetric_category.{u} (Type u) := -symmetric_of_chosen_finite_products (types.terminal_limit_cone) (types.binary_product_limit_cone) - @[simp] lemma tensor_apply {W X Y Z : Type u} (f : W ⟶ X) (g : Y ⟶ Z) (p : W ⊗ Y) : (f ⊗ g) p = (f p.1, g p.2) := rfl @@ -43,42 +41,11 @@ symmetric_of_chosen_finite_products (types.terminal_limit_cone) (types.binary_pr @[simp] lemma associator_inv_apply {X Y Z : Type u} {x : X} {y : Y} {z : Z} : ((α_ X Y Z).inv : X ⊗ (Y ⊗ Z) → (X ⊗ Y) ⊗ Z) (x, (y, z)) = ((x, y), z) := rfl -@[simp] lemma braiding_hom_apply {X Y : Type u} {x : X} {y : Y} : - ((β_ X Y).hom : X ⊗ Y → Y ⊗ X) (x, y) = (y, x) := rfl -@[simp] lemma braiding_inv_apply {X Y : Type u} {x : X} {y : Y} : - ((β_ X Y).inv : Y ⊗ X → X ⊗ Y) (y, x) = (x, y) := rfl - -open opposite - -open monoidal_category - -/-- `(𝟙_ C ⟶ -)` is a lax monoidal functor to `Type`. -/ -def coyoneda_tensor_unit (C : Type u) [category.{v} C] [monoidal_category C] : - lax_monoidal_functor C (Type v) := -{ ε := λ p, 𝟙 _, - μ := λ X Y p, (λ_ (𝟙_ C)).inv ≫ (p.1 ⊗ p.2), - μ_natural' := by tidy, - associativity' := λ X Y Z, begin - ext ⟨⟨f, g⟩, h⟩, dsimp at f g h, - dsimp, simp only [iso.cancel_iso_inv_left, category.assoc], - conv_lhs { rw [←category.id_comp h, tensor_comp, category.assoc, associator_naturality, - ←category.assoc, unitors_inv_equal, triangle_assoc_comp_right_inv], }, - conv_rhs { rw [←category.id_comp f, tensor_comp], }, - end, - left_unitality' := by tidy, - right_unitality' := λ X, begin - ext ⟨f, ⟨⟩⟩, dsimp at f, - dsimp, simp only [category.assoc], - rw [right_unitor_naturality, unitors_inv_equal, iso.inv_hom_id_assoc], - end, - ..coyoneda.obj (op (𝟙_ C)) }. - -noncomputable theory - /-- If `F` is a monoidal functor out of `Type`, it takes the (n+1)st cartesian power of a type to the image of that type, tensored with the image of the nth cartesian power. -/ -- We don't yet have an API for tensor products indexed by finite ordered types, -- but it would be nice to state how monoidal functors preserve these. +noncomputable def monoidal_functor.map_pi {C : Type*} [category C] [monoidal_category C] (F : monoidal_functor Type* C) (n : ℕ) (β : Type*) : F.obj (fin (n+1) → β) ≅ F.obj β ⊗ F.obj (fin n → β) := diff --git a/src/category_theory/monoidal/types/coyoneda.lean b/src/category_theory/monoidal/types/coyoneda.lean new file mode 100644 index 0000000000000..f9204ee86ec1c --- /dev/null +++ b/src/category_theory/monoidal/types/coyoneda.lean @@ -0,0 +1,45 @@ +/- +Copyright (c) 2018 Michael Jendrusch. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Michael Jendrusch, Scott Morrison +-/ +import category_theory.monoidal.types.basic +import category_theory.monoidal.coherence_lemmas + +/-! +# `(𝟙_ C ⟶ -)` is a lax monoidal functor to `Type` +-/ + +open category_theory +open category_theory.limits +open tactic + +universes v u + +namespace category_theory + +open opposite +open monoidal_category + +/-- `(𝟙_ C ⟶ -)` is a lax monoidal functor to `Type`. -/ +def coyoneda_tensor_unit (C : Type u) [category.{v} C] [monoidal_category C] : + lax_monoidal_functor C (Type v) := +{ ε := λ p, 𝟙 _, + μ := λ X Y p, (λ_ (𝟙_ C)).inv ≫ (p.1 ⊗ p.2), + μ_natural' := by tidy, + associativity' := λ X Y Z, begin + ext ⟨⟨f, g⟩, h⟩, dsimp at f g h, + dsimp, simp only [iso.cancel_iso_inv_left, category.assoc], + conv_lhs { rw [←category.id_comp h, tensor_comp, category.assoc, associator_naturality, + ←category.assoc, unitors_inv_equal, triangle_assoc_comp_right_inv], }, + conv_rhs { rw [←category.id_comp f, tensor_comp], }, + end, + left_unitality' := by tidy, + right_unitality' := λ X, begin + ext ⟨f, ⟨⟩⟩, dsimp at f, + dsimp, simp only [category.assoc], + rw [right_unitor_naturality, unitors_inv_equal, iso.inv_hom_id_assoc], + end, + ..coyoneda.obj (op (𝟙_ C)) }. + +end category_theory diff --git a/src/category_theory/monoidal/types/symmetric.lean b/src/category_theory/monoidal/types/symmetric.lean new file mode 100644 index 0000000000000..3aab388e36ba6 --- /dev/null +++ b/src/category_theory/monoidal/types/symmetric.lean @@ -0,0 +1,28 @@ +/- +Copyright (c) 2018 Michael Jendrusch. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Michael Jendrusch, Scott Morrison +-/ +import category_theory.monoidal.of_chosen_finite_products.symmetric +import category_theory.monoidal.types.basic + +/-! +# The category of types is a symmetric monoidal category +-/ + +open category_theory +open category_theory.limits + +universes v u + +namespace category_theory + +instance types_symmetric : symmetric_category.{u} (Type u) := +symmetric_of_chosen_finite_products (types.terminal_limit_cone) (types.binary_product_limit_cone) + +@[simp] lemma braiding_hom_apply {X Y : Type u} {x : X} {y : Y} : + ((β_ X Y).hom : X ⊗ Y → Y ⊗ X) (x, y) = (y, x) := rfl +@[simp] lemma braiding_inv_apply {X Y : Type u} {x : X} {y : Y} : + ((β_ X Y).inv : Y ⊗ X → X ⊗ Y) (y, x) = (x, y) := rfl + +end category_theory diff --git a/src/representation_theory/Action.lean b/src/representation_theory/Action.lean index 21c36f2b712a8..2a6fc140cb831 100644 --- a/src/representation_theory/Action.lean +++ b/src/representation_theory/Action.lean @@ -14,7 +14,7 @@ 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.monoidal.types +import category_theory.monoidal.types.symmetric import category_theory.abelian.functor_category import category_theory.abelian.transfer import category_theory.conj From 83a66c8775fa14ee5180c85cab98e970956401ad Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Thu, 18 May 2023 22:00:43 +0000 Subject: [PATCH 1035/1291] chore(measure_theory/special_functions/basic): split (#19040) Split out the facts about `is_R_or_C` from `measure_theory/function/special_functions/basic` (a foundational file, imported by the Bochner integral construction). These facts are heavier-weight than one would expect because the fact that an `is_R_or_C` field is a proper space currently passes through the corresponding fact for a general finite-dimensional normed space. --- src/measure_theory/function/l2_space.lean | 1 + .../function/special_functions/basic.lean | 69 +-------------- .../function/special_functions/is_R_or_C.lean | 84 +++++++++++++++++++ 3 files changed, 86 insertions(+), 68 deletions(-) create mode 100644 src/measure_theory/function/special_functions/is_R_or_C.lean diff --git a/src/measure_theory/function/l2_space.lean b/src/measure_theory/function/l2_space.lean index a46364500b128..6ab58ef0c4a86 100644 --- a/src/measure_theory/function/l2_space.lean +++ b/src/measure_theory/function/l2_space.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Rémy Degenne. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ +import data.is_R_or_C.lemmas import measure_theory.function.strongly_measurable.inner import measure_theory.integral.set_integral diff --git a/src/measure_theory/function/special_functions/basic.lean b/src/measure_theory/function/special_functions/basic.lean index 66d515bbc5706..f2a4b2b4397e5 100644 --- a/src/measure_theory/function/special_functions/basic.lean +++ b/src/measure_theory/function/special_functions/basic.lean @@ -5,7 +5,6 @@ Authors: Yury Kudryashov -/ import analysis.special_functions.pow.nnreal -import data.is_R_or_C.lemmas import measure_theory.constructions.borel_space.complex /-! @@ -77,16 +76,6 @@ measurable.ite (is_closed_le continuous_const continuous_re).measurable_set A $ end complex -namespace is_R_or_C - -variables {𝕜 : Type*} [is_R_or_C 𝕜] - -@[measurability] lemma measurable_re : measurable (re : 𝕜 → ℝ) := continuous_re.measurable - -@[measurability] lemma measurable_im : measurable (im : 𝕜 → ℝ) := continuous_im.measurable - -end is_R_or_C - section real_composition open real variables {α : Type*} {m : measurable_space α} {f : α → ℝ} (hf : measurable f) @@ -141,63 +130,6 @@ measurable_log.comp hf end complex_composition -section is_R_or_C_composition - -variables {α 𝕜 : Type*} [is_R_or_C 𝕜] {m : measurable_space α} - {f : α → 𝕜} {μ : measure_theory.measure α} - -include m - -@[measurability] lemma measurable.re (hf : measurable f) : measurable (λ x, is_R_or_C.re (f x)) := -is_R_or_C.measurable_re.comp hf - -@[measurability] lemma ae_measurable.re (hf : ae_measurable f μ) : - ae_measurable (λ x, is_R_or_C.re (f x)) μ := -is_R_or_C.measurable_re.comp_ae_measurable hf - -@[measurability] lemma measurable.im (hf : measurable f) : measurable (λ x, is_R_or_C.im (f x)) := -is_R_or_C.measurable_im.comp hf - -@[measurability] lemma ae_measurable.im (hf : ae_measurable f μ) : - ae_measurable (λ x, is_R_or_C.im (f x)) μ := -is_R_or_C.measurable_im.comp_ae_measurable hf - -omit m - -end is_R_or_C_composition - -section - -variables {α 𝕜 : Type*} [is_R_or_C 𝕜] [measurable_space α] - {f : α → 𝕜} {μ : measure_theory.measure α} - -@[measurability] lemma is_R_or_C.measurable_of_real : measurable (coe : ℝ → 𝕜) := -is_R_or_C.continuous_of_real.measurable - -lemma measurable_of_re_im - (hre : measurable (λ x, is_R_or_C.re (f x))) - (him : measurable (λ x, is_R_or_C.im (f x))) : measurable f := -begin - convert (is_R_or_C.measurable_of_real.comp hre).add - ((is_R_or_C.measurable_of_real.comp him).mul_const is_R_or_C.I), - { ext1 x, - exact (is_R_or_C.re_add_im _).symm }, - all_goals { apply_instance }, -end - -lemma ae_measurable_of_re_im - (hre : ae_measurable (λ x, is_R_or_C.re (f x)) μ) - (him : ae_measurable (λ x, is_R_or_C.im (f x)) μ) : ae_measurable f μ := -begin - convert (is_R_or_C.measurable_of_real.comp_ae_measurable hre).add - ((is_R_or_C.measurable_of_real.comp_ae_measurable him).mul_const is_R_or_C.I), - { ext1 x, - exact (is_R_or_C.re_add_im _).symm }, - all_goals { apply_instance }, -end - -end - section pow_instances instance complex.has_measurable_pow : has_measurable_pow ℂ ℂ := @@ -230,3 +162,4 @@ end pow_instances -- Guard against import creep: assert_not_exists inner_product_space assert_not_exists real.arctan +assert_not_exists finite_dimensional.proper diff --git a/src/measure_theory/function/special_functions/is_R_or_C.lean b/src/measure_theory/function/special_functions/is_R_or_C.lean new file mode 100644 index 0000000000000..5ec43cd5ae171 --- /dev/null +++ b/src/measure_theory/function/special_functions/is_R_or_C.lean @@ -0,0 +1,84 @@ +/- +Copyright (c) 2020 Yury Kudryashov. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yury Kudryashov +-/ + +import measure_theory.function.special_functions.basic +import data.is_R_or_C.lemmas + +/-! +# Measurability of the basic `is_R_or_C` functions + +-/ + +noncomputable theory +open_locale nnreal ennreal + + +namespace is_R_or_C + +variables {𝕜 : Type*} [is_R_or_C 𝕜] + +@[measurability] lemma measurable_re : measurable (re : 𝕜 → ℝ) := continuous_re.measurable + +@[measurability] lemma measurable_im : measurable (im : 𝕜 → ℝ) := continuous_im.measurable + +end is_R_or_C + +section is_R_or_C_composition + +variables {α 𝕜 : Type*} [is_R_or_C 𝕜] {m : measurable_space α} + {f : α → 𝕜} {μ : measure_theory.measure α} + +include m + +@[measurability] lemma measurable.re (hf : measurable f) : measurable (λ x, is_R_or_C.re (f x)) := +is_R_or_C.measurable_re.comp hf + +@[measurability] lemma ae_measurable.re (hf : ae_measurable f μ) : + ae_measurable (λ x, is_R_or_C.re (f x)) μ := +is_R_or_C.measurable_re.comp_ae_measurable hf + +@[measurability] lemma measurable.im (hf : measurable f) : measurable (λ x, is_R_or_C.im (f x)) := +is_R_or_C.measurable_im.comp hf + +@[measurability] lemma ae_measurable.im (hf : ae_measurable f μ) : + ae_measurable (λ x, is_R_or_C.im (f x)) μ := +is_R_or_C.measurable_im.comp_ae_measurable hf + +omit m + +end is_R_or_C_composition + +section + +variables {α 𝕜 : Type*} [is_R_or_C 𝕜] [measurable_space α] + {f : α → 𝕜} {μ : measure_theory.measure α} + +@[measurability] lemma is_R_or_C.measurable_of_real : measurable (coe : ℝ → 𝕜) := +is_R_or_C.continuous_of_real.measurable + +lemma measurable_of_re_im + (hre : measurable (λ x, is_R_or_C.re (f x))) + (him : measurable (λ x, is_R_or_C.im (f x))) : measurable f := +begin + convert (is_R_or_C.measurable_of_real.comp hre).add + ((is_R_or_C.measurable_of_real.comp him).mul_const is_R_or_C.I), + { ext1 x, + exact (is_R_or_C.re_add_im _).symm }, + all_goals { apply_instance }, +end + +lemma ae_measurable_of_re_im + (hre : ae_measurable (λ x, is_R_or_C.re (f x)) μ) + (him : ae_measurable (λ x, is_R_or_C.im (f x)) μ) : ae_measurable f μ := +begin + convert (is_R_or_C.measurable_of_real.comp_ae_measurable hre).add + ((is_R_or_C.measurable_of_real.comp_ae_measurable him).mul_const is_R_or_C.I), + { ext1 x, + exact (is_R_or_C.re_add_im _).symm }, + all_goals { apply_instance }, +end + +end From aa6669832974f87406a3d9d70fc5707a60546207 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 18 May 2023 23:40:22 +0000 Subject: [PATCH 1036/1291] feat(topology/algebra/module/star): continuity results for star_modules (#19037) Notably this adds `starL`. This also fixes some unnecessary typeclass arguments in `star_linear_equiv`. Co-authored-by: ADedecker --- src/algebra/star/module.lean | 3 +- src/analysis/normed_space/star/basic.lean | 5 ++ src/topology/algebra/module/star.lean | 69 +++++++++++++++++++++++ 3 files changed, 76 insertions(+), 1 deletion(-) create mode 100644 src/topology/algebra/module/star.lean diff --git a/src/algebra/star/module.lean b/src/algebra/star/module.lean index 17e971650b8e8..3bc5edfa0e127 100644 --- a/src/algebra/star/module.lean +++ b/src/algebra/star/module.lean @@ -63,7 +63,8 @@ end smul_lemmas then `star` is a semilinear equivalence. -/ @[simps] def star_linear_equiv (R : Type*) {A : Type*} - [comm_ring R] [star_ring R] [semiring A] [star_ring A] [module R A] [star_module R A] : + [comm_semiring R] [star_ring R] [add_comm_monoid A] [star_add_monoid A] [module R A] + [star_module R A] : A ≃ₗ⋆[R] A := { to_fun := star, map_smul' := star_smul, diff --git a/src/analysis/normed_space/star/basic.lean b/src/analysis/normed_space/star/basic.lean index b19ac63146511..722dd00fc35c3 100644 --- a/src/analysis/normed_space/star/basic.lean +++ b/src/analysis/normed_space/star/basic.lean @@ -10,6 +10,7 @@ import analysis.normed_space.linear_isometry import algebra.star.self_adjoint import algebra.star.unitary import topology.algebra.star_subalgebra +import topology.algebra.module.star /-! # Normed star rings and algebras @@ -257,6 +258,10 @@ variables {𝕜} lemma starₗᵢ_apply {x : E} : starₗᵢ 𝕜 x = star x := rfl +@[simp] lemma starₗᵢ_to_continuous_linear_equiv : + (starₗᵢ 𝕜 : E ≃ₗᵢ⋆[𝕜] E).to_continuous_linear_equiv = (starL 𝕜 : E ≃L⋆[𝕜] E) := +continuous_linear_equiv.ext rfl + end starₗᵢ namespace star_subalgebra diff --git a/src/topology/algebra/module/star.lean b/src/topology/algebra/module/star.lean new file mode 100644 index 0000000000000..314d64d6b779c --- /dev/null +++ b/src/topology/algebra/module/star.lean @@ -0,0 +1,69 @@ +/- +Copyright (c) 2023 Eric Wieser. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Eric Wieser, Frédéric Dupuis +-/ +import algebra.star.module +import topology.algebra.module.basic +import topology.algebra.star + +/-! +# The star operation, bundled as a continuous star-linear equiv +-/ + +/-- If `A` is a topological module over a commutative `R` with compatible actions, +then `star` is a continuous semilinear equivalence. -/ +@[simps] +def starL (R : Type*) {A : Type*} + [comm_semiring R] [star_ring R] [add_comm_monoid A] [star_add_monoid A] [module R A] + [star_module R A] [topological_space A] [has_continuous_star A] : + A ≃L⋆[R] A := +{ to_linear_equiv := star_linear_equiv R, + continuous_to_fun := continuous_star, + continuous_inv_fun := continuous_star } + +variables (R : Type*) (A : Type*) + [semiring R] [star_semigroup R] [has_trivial_star R] + [add_comm_group A] [module R A] [star_add_monoid A] [star_module R A] + [invertible (2 : R)] + [topological_space A] + +lemma continuous_self_adjoint_part [has_continuous_add A] [has_continuous_star A] + [has_continuous_const_smul R A] : + continuous (@self_adjoint_part R A _ _ _ _ _ _ _ _) := +((continuous_const_smul _).comp $ continuous_id.add continuous_star).subtype_mk _ + +lemma continuous_skew_adjoint_part [has_continuous_sub A] [has_continuous_star A] + [has_continuous_const_smul R A] : + continuous (@skew_adjoint_part R A _ _ _ _ _ _ _ _) := +((continuous_const_smul _).comp $ continuous_id.sub continuous_star).subtype_mk _ + +lemma continuous_decompose_prod_adjoint [topological_add_group A] [has_continuous_star A] + [has_continuous_const_smul R A] : + continuous (@star_module.decompose_prod_adjoint R A _ _ _ _ _ _ _ _) := +(continuous_self_adjoint_part R A).prod_mk (continuous_skew_adjoint_part R A) + +lemma continuous_decompose_prod_adjoint_symm [topological_add_group A] : + continuous (@star_module.decompose_prod_adjoint R A _ _ _ _ _ _ _ _).symm := +(continuous_subtype_coe.comp continuous_fst).add (continuous_subtype_coe.comp continuous_snd) + +/-- The self-adjoint part of an element of a star module, as a continuous linear map. -/ +@[simps] def self_adjoint_partL [has_continuous_add A] [has_continuous_star A] + [has_continuous_const_smul R A] : A →L[R] self_adjoint A := +{ to_linear_map := self_adjoint_part R, + cont := continuous_self_adjoint_part _ _ } + +/-- The skew-adjoint part of an element of a star module, as a continuous linear map. -/ +@[simps] def skew_adjoint_partL [has_continuous_sub A] [has_continuous_star A] + [has_continuous_const_smul R A] : A →L[R] skew_adjoint A := +{ to_linear_map := skew_adjoint_part R, + cont := continuous_skew_adjoint_part _ _ } + +/-- The decomposition of elements of a star module into their self- and skew-adjoint parts, +as a continuous linear equivalence. -/ +@[simps] def star_module.decompose_prod_adjointL [topological_add_group A] [has_continuous_star A] + [has_continuous_const_smul R A] : + A ≃L[R] self_adjoint A × skew_adjoint A := +{ to_linear_equiv := star_module.decompose_prod_adjoint R A, + continuous_to_fun := continuous_decompose_prod_adjoint _ _, + continuous_inv_fun := continuous_decompose_prod_adjoint_symm _ _ } From fd5edc43dc4f10b85abfe544b88f82cf13c5f844 Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Thu, 18 May 2023 23:40:24 +0000 Subject: [PATCH 1037/1291] chore(measure_theory/{constructions/prod,measure/lebesgue}): split (#19039) Reorganize some files in measure theory, with the goal of removing the dependence of Lebesgue measure on the Bochner integral. `measure_theory/constructions/prod` is split into `prod/basic` and `prod/integral`; the former imports the Lebesgue integral but not the Bochner integral. `measure_theory/measure/haar` is renamed `measure_theory/measure/haar/basic`. `measure_theory/measure/haar_of_inner` is renamed `measure_theory/measure/haar/inner_product_space`. `measure_theory/measure/haar_of_basis` is renamed `measure_theory/measure/haar/of_basis`. `measure_theory/measure/lebesgue` is split into `measure_theory/measure/lebesgue/basic` and `measure_theory/measure/lebesgue/integral`, with the former not importing the Bochner integral. `measure_theory/measure/complex_lebesgue` is renamed `measure_theory/measure/lebesgue/complex`. `measure_theory/measure/haar_lebesgue` is split into `measure_theory/measure/lebesgue/eq_haar` and `measure_theory/measure/haar/normed_space`, with the former not importing the Bochner integral. A few lemmas about Haar measure in normed spaces or fields are also moved from `measure_theory/group/measure` to the newly-created `measure_theory/measure/haar/normed_space`, since the former no longer imports `normed_space`. --- .../100-theorems-list/9_area_of_a_circle.lean | 3 +- counterexamples/phillips.lean | 4 +- src/analysis/box_integral/integrability.lean | 1 + .../box_integral/partition/measure.lean | 2 +- .../calculus/bump_function_findim.lean | 2 +- src/analysis/complex/cauchy_integral.lean | 2 +- src/analysis/convex/measure.lean | 2 +- src/analysis/convolution.lean | 7 +- src/analysis/fourier/fourier_transform.lean | 2 +- src/analysis/fourier/poisson_summation.lean | 1 + .../fourier/riemann_lebesgue_lemma.lean | 8 +- .../special_functions/improper_integrals.lean | 5 +- .../special_functions/japanese_bracket.lean | 2 +- src/measure_theory/constructions/pi.lean | 2 +- .../{prod.lean => prod/basic.lean} | 416 +-------------- .../constructions/prod/integral.lean | 474 ++++++++++++++++++ .../covering/besicovitch_vector_space.lean | 2 +- src/measure_theory/covering/one_dim.lean | 2 +- src/measure_theory/function/jacobian.lean | 2 +- .../group/geometry_of_numbers.lean | 2 +- src/measure_theory/group/measure.lean | 24 +- src/measure_theory/group/prod.lean | 2 +- .../integral/divergence_theorem.lean | 1 + .../integral/interval_integral.lean | 2 +- src/measure_theory/integral/periodic.lean | 4 +- .../integral/torus_integral.lean | 1 + .../measure/{haar.lean => haar/basic.lean} | 0 .../inner_product_space.lean} | 2 +- .../measure/haar/normed_space.lean | 102 ++++ .../of_basis.lean} | 2 +- .../quotient.lean} | 2 +- src/measure_theory/measure/hausdorff.lean | 2 +- .../{lebesgue.lean => lebesgue/basic.lean} | 85 +--- .../complex.lean} | 5 +- .../eq_haar.lean} | 53 +- .../measure/lebesgue/integral.lean | 97 ++++ src/measure_theory/measure/portmanteau.lean | 2 +- src/number_theory/liouville/measure.lean | 2 +- src/probability/density.lean | 2 +- src/probability/kernel/basic.lean | 4 +- test/monotonicity.lean | 3 +- 41 files changed, 738 insertions(+), 600 deletions(-) rename src/measure_theory/constructions/{prod.lean => prod/basic.lean} (61%) create mode 100644 src/measure_theory/constructions/prod/integral.lean rename src/measure_theory/measure/{haar.lean => haar/basic.lean} (100%) rename src/measure_theory/measure/{haar_of_inner.lean => haar/inner_product_space.lean} (98%) create mode 100644 src/measure_theory/measure/haar/normed_space.lean rename src/measure_theory/measure/{haar_of_basis.lean => haar/of_basis.lean} (99%) rename src/measure_theory/measure/{haar_quotient.lean => haar/quotient.lean} (99%) rename src/measure_theory/measure/{lebesgue.lean => lebesgue/basic.lean} (88%) rename src/measure_theory/measure/{complex_lebesgue.lean => lebesgue/complex.lean} (91%) rename src/measure_theory/measure/{haar_lebesgue.lean => lebesgue/eq_haar.lean} (93%) create mode 100644 src/measure_theory/measure/lebesgue/integral.lean diff --git a/archive/100-theorems-list/9_area_of_a_circle.lean b/archive/100-theorems-list/9_area_of_a_circle.lean index 9b628e2f8cc89..76a34c96e6023 100644 --- a/archive/100-theorems-list/9_area_of_a_circle.lean +++ b/archive/100-theorems-list/9_area_of_a_circle.lean @@ -3,9 +3,10 @@ Copyright (c) 2021 James Arthur, Benjamin Davidson, Andrew Souther. All rights r Released under Apache 2.0 license as described in the file LICENSE. Authors: James Arthur, Benjamin Davidson, Andrew Souther -/ -import measure_theory.integral.fund_thm_calculus import analysis.special_functions.sqrt import analysis.special_functions.trigonometric.inverse_deriv +import measure_theory.integral.fund_thm_calculus +import measure_theory.measure.lebesgue.integral /-! # Freek № 9: The Area of a Circle diff --git a/counterexamples/phillips.lean b/counterexamples/phillips.lean index d5221ec1e4383..82bdaa79e5d5d 100644 --- a/counterexamples/phillips.lean +++ b/counterexamples/phillips.lean @@ -4,7 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ import analysis.normed_space.hahn_banach.extension -import measure_theory.measure.lebesgue +import measure_theory.integral.set_integral +import measure_theory.measure.lebesgue.basic +import topology.continuous_function.bounded /-! # A counterexample on Pettis integrability diff --git a/src/analysis/box_integral/integrability.lean b/src/analysis/box_integral/integrability.lean index eba1bdd4c83d7..bc9d7862b4ea9 100644 --- a/src/analysis/box_integral/integrability.lean +++ b/src/analysis/box_integral/integrability.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import analysis.box_integral.basic +import measure_theory.integral.set_integral import measure_theory.measure.regular /-! diff --git a/src/analysis/box_integral/partition/measure.lean b/src/analysis/box_integral/partition/measure.lean index b7f96b106aefc..e61daf8b963c1 100644 --- a/src/analysis/box_integral/partition/measure.lean +++ b/src/analysis/box_integral/partition/measure.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import analysis.box_integral.partition.additive -import measure_theory.measure.lebesgue +import measure_theory.measure.lebesgue.basic /-! # Box-additive functions defined by measures diff --git a/src/analysis/calculus/bump_function_findim.lean b/src/analysis/calculus/bump_function_findim.lean index 453e2bc05af9b..67a52967d0b5a 100644 --- a/src/analysis/calculus/bump_function_findim.lean +++ b/src/analysis/calculus/bump_function_findim.lean @@ -6,7 +6,7 @@ Authors: Sébastien Gouëzel import analysis.calculus.series import analysis.convolution import analysis.inner_product_space.euclidean_dist -import measure_theory.measure.haar_lebesgue +import measure_theory.measure.haar.normed_space import data.set.pointwise.support /-! diff --git a/src/analysis/complex/cauchy_integral.lean b/src/analysis/complex/cauchy_integral.lean index d285e42c300b9..18a80345abc2e 100644 --- a/src/analysis/complex/cauchy_integral.lean +++ b/src/analysis/complex/cauchy_integral.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ -import measure_theory.measure.complex_lebesgue +import measure_theory.measure.lebesgue.complex import measure_theory.integral.divergence_theorem import measure_theory.integral.circle_integral import analysis.calculus.dslope diff --git a/src/analysis/convex/measure.lean b/src/analysis/convex/measure.lean index e2366037f7134..b4e530389b847 100644 --- a/src/analysis/convex/measure.lean +++ b/src/analysis/convex/measure.lean @@ -5,7 +5,7 @@ Authors: Yury Kudryashov -/ import analysis.convex.topology import analysis.normed_space.add_torsor_bases -import measure_theory.measure.haar_lebesgue +import measure_theory.measure.lebesgue.eq_haar /-! # Convex sets are null-measurable diff --git a/src/analysis/convolution.lean b/src/analysis/convolution.lean index c46120ed87890..8d79eeab2c1a9 100644 --- a/src/analysis/convolution.lean +++ b/src/analysis/convolution.lean @@ -3,12 +3,13 @@ Copyright (c) 2022 Floris van Doorn. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn -/ +import analysis.calculus.bump_function_inner +import analysis.calculus.parametric_integral +import measure_theory.constructions.prod.integral +import measure_theory.function.locally_integrable import measure_theory.group.integration import measure_theory.group.prod -import measure_theory.function.locally_integrable -import analysis.calculus.bump_function_inner import measure_theory.integral.interval_integral -import analysis.calculus.parametric_integral /-! # Convolution of functions diff --git a/src/analysis/fourier/fourier_transform.lean b/src/analysis/fourier/fourier_transform.lean index dc2b91809dbb5..c843af425d2a9 100644 --- a/src/analysis/fourier/fourier_transform.lean +++ b/src/analysis/fourier/fourier_transform.lean @@ -6,7 +6,7 @@ Authors: David Loeffler import analysis.complex.circle import measure_theory.group.integration -import measure_theory.measure.haar_of_basis +import measure_theory.measure.haar.of_basis /-! # The Fourier transform diff --git a/src/analysis/fourier/poisson_summation.lean b/src/analysis/fourier/poisson_summation.lean index 17e5fae7dc9cc..04cff9f36353d 100644 --- a/src/analysis/fourier/poisson_summation.lean +++ b/src/analysis/fourier/poisson_summation.lean @@ -7,6 +7,7 @@ import analysis.fourier.add_circle import analysis.fourier.fourier_transform import analysis.p_series import analysis.schwartz_space +import measure_theory.measure.lebesgue.integral /-! # Poisson's summation formula diff --git a/src/analysis/fourier/riemann_lebesgue_lemma.lean b/src/analysis/fourier/riemann_lebesgue_lemma.lean index befd529ada06b..31cd3a8acb7ce 100644 --- a/src/analysis/fourier/riemann_lebesgue_lemma.lean +++ b/src/analysis/fourier/riemann_lebesgue_lemma.lean @@ -4,12 +4,14 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: David Loeffler -/ -import measure_theory.function.continuous_map_dense -import measure_theory.group.integration import analysis.fourier.fourier_transform import analysis.inner_product_space.dual -import topology.metric_space.emetric_paracompact import analysis.inner_product_space.euclidean_dist +import measure_theory.function.continuous_map_dense +import measure_theory.group.integration +import measure_theory.integral.set_integral +import measure_theory.measure.haar.normed_space +import topology.metric_space.emetric_paracompact /-! # The Riemann-Lebesgue Lemma diff --git a/src/analysis/special_functions/improper_integrals.lean b/src/analysis/special_functions/improper_integrals.lean index 1d10e1374d9ea..fb45ccfb72a8f 100644 --- a/src/analysis/special_functions/improper_integrals.lean +++ b/src/analysis/special_functions/improper_integrals.lean @@ -3,10 +3,11 @@ Copyright (c) 2023 David Loeffler. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: David Loeffler -/ -import measure_theory.integral.integral_eq_improper +import analysis.special_functions.integrals import measure_theory.group.integration import measure_theory.integral.exp_decay -import analysis.special_functions.integrals +import measure_theory.integral.integral_eq_improper +import measure_theory.measure.lebesgue.integral /-! # Evaluation of specific improper integrals diff --git a/src/analysis/special_functions/japanese_bracket.lean b/src/analysis/special_functions/japanese_bracket.lean index d287916d88f20..94cf135552127 100644 --- a/src/analysis/special_functions/japanese_bracket.lean +++ b/src/analysis/special_functions/japanese_bracket.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Moritz Doll. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Moritz Doll -/ -import measure_theory.measure.haar_lebesgue +import measure_theory.measure.lebesgue.eq_haar import measure_theory.integral.layercake /-! diff --git a/src/measure_theory/constructions/pi.lean b/src/measure_theory/constructions/pi.lean index 4fb749a467fee..2b2c400417f35 100644 --- a/src/measure_theory/constructions/pi.lean +++ b/src/measure_theory/constructions/pi.lean @@ -3,7 +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 -/ -import measure_theory.constructions.prod +import measure_theory.constructions.prod.basic import measure_theory.group.measure import topology.constructions diff --git a/src/measure_theory/constructions/prod.lean b/src/measure_theory/constructions/prod/basic.lean similarity index 61% rename from src/measure_theory/constructions/prod.lean rename to src/measure_theory/constructions/prod/basic.lean index 642c354a80fb6..f2afcf79db87c 100644 --- a/src/measure_theory/constructions/prod.lean +++ b/src/measure_theory/constructions/prod/basic.lean @@ -5,7 +5,7 @@ Authors: Floris van Doorn -/ import measure_theory.measure.giry_monad import dynamics.ergodic.measure_preserving -import measure_theory.integral.set_integral +import measure_theory.integral.lebesgue import measure_theory.measure.open_pos /-! @@ -17,7 +17,7 @@ satisfies `(μ.prod ν) s = ∫⁻ x, ν {y | (x, y) ∈ s} ∂μ`. We also have `(μ.prod ν) (s ×ˢ t) = μ s * ν t`, i.e. the measure of a rectangle is the product of the measures of the sides. -We also prove Tonelli's theorem and Fubini's theorem. +We also prove Tonelli's theorem. ## Main definition @@ -35,14 +35,6 @@ We also prove Tonelli's theorem and Fubini's theorem. a variant with `_symm` appended, where the order of integration is reversed. The lemma `measurable.lintegral_prod_right'` states that the inner integral of the right-hand side is measurable. -* `measure_theory.integrable_prod_iff` states that a binary function is integrable iff both - * `y ↦ f (x, y)` is integrable for almost every `x`, and - * the function `x ↦ ∫ ‖f (x, y)‖ dy` is integrable. -* `measure_theory.integral_prod`: Fubini's theorem. It states that for a integrable function - `α × β → E` (where `E` is a second countable Banach space) we have - `∫ z, f z ∂(μ.prod ν) = ∫ x, ∫ y, f (x, y) ∂ν ∂μ`. This theorem has the same variants as - Tonelli's theorem. The lemma `measure_theory.integrable.integral_prod_right` states that the - inner integral of the right-hand side is integrable. ## Implementation Notes @@ -51,12 +43,12 @@ functions in uncurried form (`α × β → γ`). The former often has an assumpt `measurable (uncurry f)`, which could be inconvenient to discharge, but for the latter it is more common that the function has to be given explicitly, since Lean cannot synthesize the function by itself. We name the lemmas about the uncurried form with a prime. -Tonelli's theorem and Fubini's theorem have a different naming scheme, since the version for the -uncurried version is reversed. +Tonelli's theorem has a different naming scheme, since the version for the uncurried version is +reversed. ## Tags -product measure, Fubini's theorem, Tonelli's theorem, Fubini-Tonelli theorem +product measure, Tonelli's theorem, Fubini-Tonelli theorem -/ noncomputable theory @@ -236,81 +228,6 @@ lemma measurable.lintegral_prod_left [sigma_finite μ] {f : α → β → ℝ≥ (hf : measurable (uncurry f)) : measurable (λ y, ∫⁻ x, f x y ∂μ) := hf.lintegral_prod_left' -lemma measurable_set_integrable [sigma_finite ν] ⦃f : α → β → E⦄ - (hf : strongly_measurable (uncurry f)) : measurable_set {x | integrable (f x) ν} := -begin - simp_rw [integrable, hf.of_uncurry_left.ae_strongly_measurable, true_and], - exact measurable_set_lt (measurable.lintegral_prod_right hf.ennnorm) measurable_const -end - -section -variables [normed_space ℝ E] [complete_space E] - -/-- The Bochner integral is measurable. This shows that the integrand of (the right-hand-side of) - Fubini's theorem is measurable. - This version has `f` in curried form. -/ -lemma measure_theory.strongly_measurable.integral_prod_right [sigma_finite ν] ⦃f : α → β → E⦄ - (hf : strongly_measurable (uncurry f)) : strongly_measurable (λ x, ∫ y, f x y ∂ν) := -begin - borelize E, - haveI : separable_space (range (uncurry f) ∪ {0} : set E) := - hf.separable_space_range_union_singleton, - let s : ℕ → simple_func (α × β) E := simple_func.approx_on _ hf.measurable - (range (uncurry f) ∪ {0}) 0 (by simp), - let s' : ℕ → α → simple_func β E := λ n x, (s n).comp (prod.mk x) measurable_prod_mk_left, - let f' : ℕ → α → E := λ n, {x | integrable (f x) ν}.indicator - (λ x, (s' n x).integral ν), - have hf' : ∀ n, strongly_measurable (f' n), - { intro n, refine strongly_measurable.indicator _ (measurable_set_integrable hf), - have : ∀ x, (s' n x).range.filter (λ x, x ≠ 0) ⊆ (s n).range, - { intros x, refine finset.subset.trans (finset.filter_subset _ _) _, intro y, - simp_rw [simple_func.mem_range], rintro ⟨z, rfl⟩, exact ⟨(x, z), rfl⟩ }, - simp only [simple_func.integral_eq_sum_of_subset (this _)], - refine finset.strongly_measurable_sum _ (λ x _, _), - refine (measurable.ennreal_to_real _).strongly_measurable.smul_const _, - simp only [simple_func.coe_comp, preimage_comp] {single_pass := tt}, - apply measurable_measure_prod_mk_left, - exact (s n).measurable_set_fiber x }, - have h2f' : tendsto f' at_top (𝓝 (λ (x : α), ∫ (y : β), f x y ∂ν)), - { rw [tendsto_pi_nhds], intro x, - by_cases hfx : integrable (f x) ν, - { have : ∀ n, integrable (s' n x) ν, - { intro n, apply (hfx.norm.add hfx.norm).mono' (s' n x).ae_strongly_measurable, - apply eventually_of_forall, intro y, - simp_rw [s', simple_func.coe_comp], exact simple_func.norm_approx_on_zero_le _ _ (x, y) n }, - simp only [f', hfx, simple_func.integral_eq_integral _ (this _), indicator_of_mem, - mem_set_of_eq], - refine tendsto_integral_of_dominated_convergence (λ y, ‖f x y‖ + ‖f x y‖) - (λ n, (s' n x).ae_strongly_measurable) (hfx.norm.add hfx.norm) _ _, - { exact λ n, eventually_of_forall (λ y, simple_func.norm_approx_on_zero_le _ _ (x, y) n) }, - { refine eventually_of_forall (λ y, simple_func.tendsto_approx_on _ _ _), - apply subset_closure, - simp [-uncurry_apply_pair], } }, - { simp [f', hfx, integral_undef], } }, - exact strongly_measurable_of_tendsto _ hf' h2f' -end - -/-- The Bochner integral is measurable. This shows that the integrand of (the right-hand-side of) - Fubini's theorem is measurable. -/ -lemma measure_theory.strongly_measurable.integral_prod_right' [sigma_finite ν] ⦃f : α × β → E⦄ - (hf : strongly_measurable f) : strongly_measurable (λ x, ∫ y, f (x, y) ∂ν) := -by { rw [← uncurry_curry f] at hf, exact hf.integral_prod_right } - -/-- The Bochner integral is measurable. This shows that the integrand of (the right-hand-side of) - the symmetric version of Fubini's theorem is measurable. - This version has `f` in curried form. -/ -lemma measure_theory.strongly_measurable.integral_prod_left [sigma_finite μ] ⦃f : α → β → E⦄ - (hf : strongly_measurable (uncurry f)) : strongly_measurable (λ y, ∫ x, f x y ∂μ) := -(hf.comp_measurable measurable_swap).integral_prod_right' - -/-- The Bochner integral is measurable. This shows that the integrand of (the right-hand-side of) - the symmetric version of Fubini's theorem is measurable. -/ -lemma measure_theory.strongly_measurable.integral_prod_left' [sigma_finite μ] ⦃f : α × β → E⦄ - (hf : strongly_measurable f) : strongly_measurable (λ y, ∫ x, f (x, y) ∂μ) := -(hf.comp_measurable measurable_swap).integral_prod_right' - -end - /-! ### The product measure -/ namespace measure_theory @@ -418,18 +335,6 @@ lemma ae_measure_lt_top {s : set (α × β)} (hs : measurable_set s) (h2s : (μ.prod ν) s ≠ ∞) : ∀ᵐ x ∂μ, ν (prod.mk x ⁻¹' s) < ∞ := by { simp_rw [prod_apply hs] at h2s, refine ae_lt_top (measurable_measure_prod_mk_left hs) h2s } -lemma integrable_measure_prod_mk_left {s : set (α × β)} - (hs : measurable_set s) (h2s : (μ.prod ν) s ≠ ∞) : - integrable (λ x, (ν (prod.mk x ⁻¹' s)).to_real) μ := -begin - refine ⟨(measurable_measure_prod_mk_left hs).ennreal_to_real.ae_measurable.ae_strongly_measurable, - _⟩, - simp_rw [has_finite_integral, ennnorm_eq_of_real to_real_nonneg], - convert h2s.lt_top using 1, simp_rw [prod_apply hs], apply lintegral_congr_ae, - refine (ae_measure_lt_top hs h2s).mp _, apply eventually_of_forall, intros x hx, - rw [lt_top_iff_ne_top] at hx, simp [of_real_to_real, hx], -end - /-- Note: the assumption `hs` cannot be dropped. For a counterexample, see Walter Rudin *Real and Complex Analysis*, example (c) in section 8.9. -/ lemma measure_prod_null {s : set (α × β)} @@ -700,12 +605,6 @@ lemma ae_measurable.prod_swap [sigma_finite μ] [sigma_finite ν] {f : β × α (hf : ae_measurable f (ν.prod μ)) : ae_measurable (λ (z : α × β), f z.swap) (μ.prod ν) := by { rw ← prod_swap at hf, exact hf.comp_measurable measurable_swap } -lemma measure_theory.ae_strongly_measurable.prod_swap - {γ : Type*} [topological_space γ] [sigma_finite μ] [sigma_finite ν] {f : β × α → γ} - (hf : ae_strongly_measurable f (ν.prod μ)) : - ae_strongly_measurable (λ (z : α × β), f z.swap) (μ.prod ν) := -by { rw ← prod_swap at hf, exact hf.comp_measurable measurable_swap } - lemma ae_measurable.fst [sigma_finite ν] {f : α → γ} (hf : ae_measurable f μ) : ae_measurable (λ (z : α × β), f z.1) (μ.prod ν) := hf.comp_quasi_measure_preserving quasi_measure_preserving_fst @@ -714,33 +613,6 @@ lemma ae_measurable.snd [sigma_finite ν] {f : β → γ} (hf : ae_measurable f ν) : ae_measurable (λ (z : α × β), f z.2) (μ.prod ν) := hf.comp_quasi_measure_preserving quasi_measure_preserving_snd -lemma measure_theory.ae_strongly_measurable.fst {γ} [topological_space γ] [sigma_finite ν] - {f : α → γ} (hf : ae_strongly_measurable f μ) : - ae_strongly_measurable (λ (z : α × β), f z.1) (μ.prod ν) := -hf.comp_quasi_measure_preserving quasi_measure_preserving_fst - -lemma measure_theory.ae_strongly_measurable.snd {γ} [topological_space γ] [sigma_finite ν] - {f : β → γ} (hf : ae_strongly_measurable f ν) : - ae_strongly_measurable (λ (z : α × β), f z.2) (μ.prod ν) := -hf.comp_quasi_measure_preserving quasi_measure_preserving_snd - -/-- The Bochner integral is a.e.-measurable. - This shows that the integrand of (the right-hand-side of) Fubini's theorem is a.e.-measurable. -/ -lemma measure_theory.ae_strongly_measurable.integral_prod_right' [sigma_finite ν] - [normed_space ℝ E] [complete_space E] - ⦃f : α × β → E⦄ (hf : ae_strongly_measurable f (μ.prod ν)) : - ae_strongly_measurable (λ x, ∫ y, f (x, y) ∂ν) μ := -⟨λ x, ∫ y, hf.mk f (x, y) ∂ν, hf.strongly_measurable_mk.integral_prod_right', - by { filter_upwards [ae_ae_of_ae_prod hf.ae_eq_mk] with _ hx using integral_congr_ae hx }⟩ - -lemma measure_theory.ae_strongly_measurable.prod_mk_left - {γ : Type*} [sigma_finite ν] [topological_space γ] {f : α × β → γ} - (hf : ae_strongly_measurable f (μ.prod ν)) : ∀ᵐ x ∂μ, ae_strongly_measurable (λ y, f (x, y)) ν := -begin - filter_upwards [ae_ae_of_ae_prod hf.ae_eq_mk] with x hx, - exact ⟨λ y, hf.mk f (x, y), hf.strongly_measurable_mk.comp_measurable measurable_prod_mk_left, hx⟩ -end - end namespace measure_theory @@ -824,284 +696,6 @@ lemma lintegral_prod_mul {f : α → ℝ≥0∞} {g : β → ℝ≥0∞} ∫⁻ z, f z.1 * g z.2 ∂(μ.prod ν) = ∫⁻ x, f x ∂μ * ∫⁻ y, g y ∂ν := by simp [lintegral_prod _ (hf.fst.mul hg.snd), lintegral_lintegral_mul hf hg] -/-! ### Integrability on a product -/ -section - -lemma integrable.swap [sigma_finite μ] ⦃f : α × β → E⦄ - (hf : integrable f (μ.prod ν)) : integrable (f ∘ prod.swap) (ν.prod μ) := -⟨hf.ae_strongly_measurable.prod_swap, - (lintegral_prod_swap _ hf.ae_strongly_measurable.ennnorm : _).le.trans_lt hf.has_finite_integral⟩ - -lemma integrable_swap_iff [sigma_finite μ] ⦃f : α × β → E⦄ : - integrable (f ∘ prod.swap) (ν.prod μ) ↔ integrable f (μ.prod ν) := -⟨λ hf, by { convert hf.swap, ext ⟨x, y⟩, refl }, λ hf, hf.swap⟩ - -lemma has_finite_integral_prod_iff ⦃f : α × β → E⦄ (h1f : strongly_measurable f) : - has_finite_integral f (μ.prod ν) ↔ (∀ᵐ x ∂ μ, has_finite_integral (λ y, f (x, y)) ν) ∧ - has_finite_integral (λ x, ∫ y, ‖f (x, y)‖ ∂ν) μ := -begin - simp only [has_finite_integral, lintegral_prod_of_measurable _ h1f.ennnorm], - have : ∀ x, ∀ᵐ y ∂ν, 0 ≤ ‖f (x, y)‖ := λ x, eventually_of_forall (λ y, norm_nonneg _), - simp_rw [integral_eq_lintegral_of_nonneg_ae (this _) - (h1f.norm.comp_measurable measurable_prod_mk_left).ae_strongly_measurable, - ennnorm_eq_of_real to_real_nonneg, of_real_norm_eq_coe_nnnorm], - -- this fact is probably too specialized to be its own lemma - have : ∀ {p q r : Prop} (h1 : r → p), (r ↔ p ∧ q) ↔ (p → (r ↔ q)) := - λ p q r h1, by rw [← and.congr_right_iff, and_iff_right_of_imp h1], - rw [this], - { intro h2f, rw lintegral_congr_ae, - refine h2f.mp _, apply eventually_of_forall, intros x hx, dsimp only, - rw [of_real_to_real], rw [← lt_top_iff_ne_top], exact hx }, - { intro h2f, refine ae_lt_top _ h2f.ne, exact h1f.ennnorm.lintegral_prod_right' }, -end - -lemma has_finite_integral_prod_iff' ⦃f : α × β → E⦄ (h1f : ae_strongly_measurable f (μ.prod ν)) : - has_finite_integral f (μ.prod ν) ↔ (∀ᵐ x ∂ μ, has_finite_integral (λ y, f (x, y)) ν) ∧ - has_finite_integral (λ x, ∫ y, ‖f (x, y)‖ ∂ν) μ := -begin - rw [has_finite_integral_congr h1f.ae_eq_mk, - has_finite_integral_prod_iff h1f.strongly_measurable_mk], - apply and_congr, - { apply eventually_congr, - filter_upwards [ae_ae_of_ae_prod h1f.ae_eq_mk.symm], - assume x hx, - exact has_finite_integral_congr hx }, - { apply has_finite_integral_congr, - filter_upwards [ae_ae_of_ae_prod h1f.ae_eq_mk.symm] with _ hx - using integral_congr_ae (eventually_eq.fun_comp hx _), }, - { apply_instance, }, -end - -/-- A binary function is integrable if the function `y ↦ f (x, y)` is integrable for almost every - `x` and the function `x ↦ ∫ ‖f (x, y)‖ dy` is integrable. -/ -lemma integrable_prod_iff ⦃f : α × β → E⦄ (h1f : ae_strongly_measurable f (μ.prod ν)) : - integrable f (μ.prod ν) ↔ - (∀ᵐ x ∂ μ, integrable (λ y, f (x, y)) ν) ∧ integrable (λ x, ∫ y, ‖f (x, y)‖ ∂ν) μ := -by simp [integrable, h1f, has_finite_integral_prod_iff', h1f.norm.integral_prod_right', - h1f.prod_mk_left] - -/-- A binary function is integrable if the function `x ↦ f (x, y)` is integrable for almost every - `y` and the function `y ↦ ∫ ‖f (x, y)‖ dx` is integrable. -/ -lemma integrable_prod_iff' [sigma_finite μ] ⦃f : α × β → E⦄ - (h1f : ae_strongly_measurable f (μ.prod ν)) : - integrable f (μ.prod ν) ↔ - (∀ᵐ y ∂ ν, integrable (λ x, f (x, y)) μ) ∧ integrable (λ y, ∫ x, ‖f (x, y)‖ ∂μ) ν := -by { convert integrable_prod_iff (h1f.prod_swap) using 1, rw [integrable_swap_iff] } - -lemma integrable.prod_left_ae [sigma_finite μ] ⦃f : α × β → E⦄ - (hf : integrable f (μ.prod ν)) : ∀ᵐ y ∂ ν, integrable (λ x, f (x, y)) μ := -((integrable_prod_iff' hf.ae_strongly_measurable).mp hf).1 - -lemma integrable.prod_right_ae [sigma_finite μ] ⦃f : α × β → E⦄ - (hf : integrable f (μ.prod ν)) : ∀ᵐ x ∂ μ, integrable (λ y, f (x, y)) ν := -hf.swap.prod_left_ae - -lemma integrable.integral_norm_prod_left ⦃f : α × β → E⦄ - (hf : integrable f (μ.prod ν)) : integrable (λ x, ∫ y, ‖f (x, y)‖ ∂ν) μ := -((integrable_prod_iff hf.ae_strongly_measurable).mp hf).2 - -lemma integrable.integral_norm_prod_right [sigma_finite μ] ⦃f : α × β → E⦄ - (hf : integrable f (μ.prod ν)) : integrable (λ y, ∫ x, ‖f (x, y)‖ ∂μ) ν := -hf.swap.integral_norm_prod_left - -lemma integrable_prod_mul {L : Type*} [is_R_or_C L] - {f : α → L} {g : β → L} (hf : integrable f μ) (hg : integrable g ν) : - integrable (λ (z : α × β), f z.1 * g z.2) (μ.prod ν) := -begin - refine (integrable_prod_iff _).2 ⟨_, _⟩, - { exact hf.1.fst.mul hg.1.snd }, - { exact eventually_of_forall (λ x, hg.const_mul (f x)) }, - { simpa only [norm_mul, integral_mul_left] using hf.norm.mul_const _ } -end - -end - -variables [normed_space ℝ E] [complete_space E] - -lemma integrable.integral_prod_left ⦃f : α × β → E⦄ - (hf : integrable f (μ.prod ν)) : integrable (λ x, ∫ y, f (x, y) ∂ν) μ := -integrable.mono hf.integral_norm_prod_left hf.ae_strongly_measurable.integral_prod_right' $ - eventually_of_forall $ λ x, (norm_integral_le_integral_norm _).trans_eq $ - (norm_of_nonneg $ integral_nonneg_of_ae $ eventually_of_forall $ - λ y, (norm_nonneg (f (x, y)) : _)).symm - -lemma integrable.integral_prod_right [sigma_finite μ] ⦃f : α × β → E⦄ - (hf : integrable f (μ.prod ν)) : integrable (λ y, ∫ x, f (x, y) ∂μ) ν := -hf.swap.integral_prod_left - -/-! ### The Bochner integral on a product -/ - -variables [sigma_finite μ] - -lemma integral_prod_swap (f : α × β → E) - (hf : ae_strongly_measurable f (μ.prod ν)) : ∫ z, f z.swap ∂(ν.prod μ) = ∫ z, f z ∂(μ.prod ν) := -begin - rw ← prod_swap at hf, - rw [← integral_map measurable_swap.ae_measurable hf, prod_swap] -end - -variables {E' : Type*} [normed_add_comm_group E'] [complete_space E'] [normed_space ℝ E'] - -/-! Some rules about the sum/difference of double integrals. They follow from `integral_add`, but - we separate them out as separate lemmas, because they involve quite some steps. -/ - -/-- Integrals commute with addition inside another integral. `F` can be any function. -/ -lemma integral_fn_integral_add ⦃f g : α × β → E⦄ (F : E → E') - (hf : integrable f (μ.prod ν)) (hg : integrable g (μ.prod ν)) : - ∫ x, F (∫ y, f (x, y) + g (x, y) ∂ν) ∂μ = ∫ x, F (∫ y, f (x, y) ∂ν + ∫ y, g (x, y) ∂ν) ∂μ := -begin - refine integral_congr_ae _, - filter_upwards [hf.prod_right_ae, hg.prod_right_ae] with _ h2f h2g, - simp [integral_add h2f h2g], -end - -/-- Integrals commute with subtraction inside another integral. - `F` can be any measurable function. -/ -lemma integral_fn_integral_sub ⦃f g : α × β → E⦄ (F : E → E') - (hf : integrable f (μ.prod ν)) (hg : integrable g (μ.prod ν)) : - ∫ x, F (∫ y, f (x, y) - g (x, y) ∂ν) ∂μ = ∫ x, F (∫ y, f (x, y) ∂ν - ∫ y, g (x, y) ∂ν) ∂μ := -begin - refine integral_congr_ae _, - filter_upwards [hf.prod_right_ae, hg.prod_right_ae] with _ h2f h2g, - simp [integral_sub h2f h2g], -end - -/-- Integrals commute with subtraction inside a lower Lebesgue integral. - `F` can be any function. -/ -lemma lintegral_fn_integral_sub ⦃f g : α × β → E⦄ - (F : E → ℝ≥0∞) (hf : integrable f (μ.prod ν)) (hg : integrable g (μ.prod ν)) : - ∫⁻ x, F (∫ y, f (x, y) - g (x, y) ∂ν) ∂μ = ∫⁻ x, F (∫ y, f (x, y) ∂ν - ∫ y, g (x, y) ∂ν) ∂μ := -begin - refine lintegral_congr_ae _, - filter_upwards [hf.prod_right_ae, hg.prod_right_ae] with _ h2f h2g, - simp [integral_sub h2f h2g], -end - -/-- Double integrals commute with addition. -/ -lemma integral_integral_add ⦃f g : α × β → E⦄ - (hf : integrable f (μ.prod ν)) (hg : integrable g (μ.prod ν)) : - ∫ x, ∫ y, f (x, y) + g (x, y) ∂ν ∂μ = ∫ x, ∫ y, f (x, y) ∂ν ∂μ + ∫ x, ∫ y, g (x, y) ∂ν ∂μ := -(integral_fn_integral_add id hf hg).trans $ - integral_add hf.integral_prod_left hg.integral_prod_left - -/-- Double integrals commute with addition. This is the version with `(f + g) (x, y)` - (instead of `f (x, y) + g (x, y)`) in the LHS. -/ -lemma integral_integral_add' ⦃f g : α × β → E⦄ - (hf : integrable f (μ.prod ν)) (hg : integrable g (μ.prod ν)) : - ∫ x, ∫ y, (f + g) (x, y) ∂ν ∂μ = ∫ x, ∫ y, f (x, y) ∂ν ∂μ + ∫ x, ∫ y, g (x, y) ∂ν ∂μ := -integral_integral_add hf hg - -/-- Double integrals commute with subtraction. -/ -lemma integral_integral_sub ⦃f g : α × β → E⦄ - (hf : integrable f (μ.prod ν)) (hg : integrable g (μ.prod ν)) : - ∫ x, ∫ y, f (x, y) - g (x, y) ∂ν ∂μ = ∫ x, ∫ y, f (x, y) ∂ν ∂μ - ∫ x, ∫ y, g (x, y) ∂ν ∂μ := -(integral_fn_integral_sub id hf hg).trans $ - integral_sub hf.integral_prod_left hg.integral_prod_left - -/-- Double integrals commute with subtraction. This is the version with `(f - g) (x, y)` - (instead of `f (x, y) - g (x, y)`) in the LHS. -/ -lemma integral_integral_sub' ⦃f g : α × β → E⦄ - (hf : integrable f (μ.prod ν)) (hg : integrable g (μ.prod ν)) : - ∫ x, ∫ y, (f - g) (x, y) ∂ν ∂μ = ∫ x, ∫ y, f (x, y) ∂ν ∂μ - ∫ x, ∫ y, g (x, y) ∂ν ∂μ := -integral_integral_sub hf hg - -/-- The map that sends an L¹-function `f : α × β → E` to `∫∫f` is continuous. -/ -lemma continuous_integral_integral : - continuous (λ (f : α × β →₁[μ.prod ν] E), ∫ x, ∫ y, f (x, y) ∂ν ∂μ) := -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, (‖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, ‖i (x, y) - g (x, y)‖₊ ∂ν ∂μ }, - swap, { exact λ i, lintegral_mono (λ x, ennnorm_integral_le_lintegral_ennnorm _) }, - show tendsto (λ (i : α × β →₁[μ.prod ν] E), - ∫⁻ 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], - refine (continuous_of_real.tendsto 0).comp _, - rw [← tendsto_iff_norm_tendsto_zero], exact tendsto_id -end - -/-- **Fubini's Theorem**: For integrable functions on `α × β`, - the Bochner integral of `f` is equal to the iterated Bochner integral. - `integrable_prod_iff` can be useful to show that the function in question in integrable. - `measure_theory.integrable.integral_prod_right` is useful to show that the inner integral - of the right-hand side is integrable. -/ -lemma integral_prod : ∀ (f : α × β → E) (hf : integrable f (μ.prod ν)), - ∫ z, f z ∂(μ.prod ν) = ∫ x, ∫ y, f (x, y) ∂ν ∂μ := -begin - apply integrable.induction, - { intros c s hs h2s, - simp_rw [integral_indicator hs, ← indicator_comp_right, - function.comp, integral_indicator (measurable_prod_mk_left hs), - set_integral_const, integral_smul_const, - integral_to_real (measurable_measure_prod_mk_left hs).ae_measurable - (ae_measure_lt_top hs h2s.ne), prod_apply hs] }, - { intros f g hfg i_f i_g hf hg, - simp_rw [integral_add' i_f i_g, integral_integral_add' i_f i_g, hf, hg] }, - { exact is_closed_eq continuous_integral continuous_integral_integral }, - { intros f g hfg i_f hf, convert hf using 1, - { exact integral_congr_ae hfg.symm }, - { refine integral_congr_ae _, - refine (ae_ae_of_ae_prod hfg).mp _, - apply eventually_of_forall, intros x hfgx, - exact integral_congr_ae (ae_eq_symm hfgx) } } -end - -/-- Symmetric version of **Fubini's Theorem**: For integrable functions on `α × β`, - the Bochner integral of `f` is equal to the iterated Bochner integral. - This version has the integrals on the right-hand side in the other order. -/ -lemma integral_prod_symm (f : α × β → E) (hf : integrable f (μ.prod ν)) : - ∫ z, f z ∂(μ.prod ν) = ∫ y, ∫ x, f (x, y) ∂μ ∂ν := -by { simp_rw [← integral_prod_swap f hf.ae_strongly_measurable], exact integral_prod _ hf.swap } - -/-- Reversed version of **Fubini's Theorem**. -/ -lemma integral_integral {f : α → β → E} (hf : integrable (uncurry f) (μ.prod ν)) : - ∫ x, ∫ y, f x y ∂ν ∂μ = ∫ z, f z.1 z.2 ∂(μ.prod ν) := -(integral_prod _ hf).symm - -/-- Reversed version of **Fubini's Theorem** (symmetric version). -/ -lemma integral_integral_symm {f : α → β → E} (hf : integrable (uncurry f) (μ.prod ν)) : - ∫ x, ∫ y, f x y ∂ν ∂μ = ∫ z, f z.2 z.1 ∂(ν.prod μ) := -(integral_prod_symm _ hf.swap).symm - -/-- Change the order of Bochner integration. -/ -lemma integral_integral_swap ⦃f : α → β → E⦄ (hf : integrable (uncurry f) (μ.prod ν)) : - ∫ x, ∫ y, f x y ∂ν ∂μ = ∫ y, ∫ x, f x y ∂μ ∂ν := -(integral_integral hf).trans (integral_prod_symm _ hf) - -/-- **Fubini's Theorem** for set integrals. -/ -lemma set_integral_prod (f : α × β → E) {s : set α} {t : set β} - (hf : integrable_on f (s ×ˢ t) (μ.prod ν)) : - ∫ z in s ×ˢ t, f z ∂(μ.prod ν) = ∫ x in s, ∫ y in t, f (x, y) ∂ν ∂μ := -begin - simp only [← measure.prod_restrict s t, integrable_on] at hf ⊢, - exact integral_prod f hf -end - -lemma integral_prod_mul {L : Type*} [is_R_or_C L] (f : α → L) (g : β → L) : - ∫ z, f z.1 * g z.2 ∂(μ.prod ν) = (∫ x, f x ∂μ) * (∫ y, g y ∂ν) := -begin - by_cases h : integrable (λ (z : α × β), f z.1 * g z.2) (μ.prod ν), - { rw integral_prod _ h, - simp_rw [integral_mul_left, integral_mul_right] }, - have H : ¬(integrable f μ) ∨ ¬(integrable g ν), - { contrapose! h, - exact integrable_prod_mul h.1 h.2 }, - cases H; - simp [integral_undef h, integral_undef H], -end - -lemma set_integral_prod_mul {L : Type*} [is_R_or_C L] - (f : α → L) (g : β → L) (s : set α) (t : set β) : - ∫ z in s ×ˢ t, f z.1 * g z.2 ∂(μ.prod ν) = (∫ x in s, f x ∂μ) * (∫ y in t, g y ∂ν) := -by simp only [← measure.prod_restrict s t, integrable_on, integral_prod_mul] - /-! ### Marginals of a measure defined on a product -/ namespace measure diff --git a/src/measure_theory/constructions/prod/integral.lean b/src/measure_theory/constructions/prod/integral.lean new file mode 100644 index 0000000000000..dab3b0377a83b --- /dev/null +++ b/src/measure_theory/constructions/prod/integral.lean @@ -0,0 +1,474 @@ +/- +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 +-/ +import measure_theory.constructions.prod.basic +import measure_theory.integral.set_integral + +/-! +# Integration with respect to the product measure + +In this file we prove Fubini's theorem. + +## Main results + +* `measure_theory.integrable_prod_iff` states that a binary function is integrable iff both + * `y ↦ f (x, y)` is integrable for almost every `x`, and + * the function `x ↦ ∫ ‖f (x, y)‖ dy` is integrable. +* `measure_theory.integral_prod`: Fubini's theorem. It states that for a integrable function + `α × β → E` (where `E` is a second countable Banach space) we have + `∫ z, f z ∂(μ.prod ν) = ∫ x, ∫ y, f (x, y) ∂ν ∂μ`. This theorem has the same variants as + Tonelli's theorem (see `measure_theory.lintegral_prod`). The lemma + `measure_theory.integrable.integral_prod_right` states that the inner integral of the right-hand + side is integrable. + +## Tags + +product measure, Fubini's theorem, Fubini-Tonelli theorem +-/ + +noncomputable theory +open_locale classical topology ennreal measure_theory +open set function real ennreal +open measure_theory measurable_space measure_theory.measure +open topological_space +open filter (hiding prod_eq map) + +variables {α α' β β' γ E : Type*} + +variables [measurable_space α] [measurable_space α'] [measurable_space β] [measurable_space β'] +variables [measurable_space γ] +variables {μ μ' : measure α} {ν ν' : measure β} {τ : measure γ} +variables [normed_add_comm_group E] + +/-! ### Measurability + +Before we define the product measure, we can talk about the measurability of operations on binary +functions. We show that if `f` is a binary measurable function, then the function that integrates +along one of the variables (using either the Lebesgue or Bochner integral) is measurable. +-/ + +lemma measurable_set_integrable [sigma_finite ν] ⦃f : α → β → E⦄ + (hf : strongly_measurable (uncurry f)) : measurable_set {x | integrable (f x) ν} := +begin + simp_rw [integrable, hf.of_uncurry_left.ae_strongly_measurable, true_and], + exact measurable_set_lt (measurable.lintegral_prod_right hf.ennnorm) measurable_const +end + +section +variables [normed_space ℝ E] [complete_space E] + +/-- The Bochner integral is measurable. This shows that the integrand of (the right-hand-side of) + Fubini's theorem is measurable. + This version has `f` in curried form. -/ +lemma measure_theory.strongly_measurable.integral_prod_right [sigma_finite ν] ⦃f : α → β → E⦄ + (hf : strongly_measurable (uncurry f)) : strongly_measurable (λ x, ∫ y, f x y ∂ν) := +begin + borelize E, + haveI : separable_space (range (uncurry f) ∪ {0} : set E) := + hf.separable_space_range_union_singleton, + let s : ℕ → simple_func (α × β) E := simple_func.approx_on _ hf.measurable + (range (uncurry f) ∪ {0}) 0 (by simp), + let s' : ℕ → α → simple_func β E := λ n x, (s n).comp (prod.mk x) measurable_prod_mk_left, + let f' : ℕ → α → E := λ n, {x | integrable (f x) ν}.indicator + (λ x, (s' n x).integral ν), + have hf' : ∀ n, strongly_measurable (f' n), + { intro n, refine strongly_measurable.indicator _ (measurable_set_integrable hf), + have : ∀ x, (s' n x).range.filter (λ x, x ≠ 0) ⊆ (s n).range, + { intros x, refine finset.subset.trans (finset.filter_subset _ _) _, intro y, + simp_rw [simple_func.mem_range], rintro ⟨z, rfl⟩, exact ⟨(x, z), rfl⟩ }, + simp only [simple_func.integral_eq_sum_of_subset (this _)], + refine finset.strongly_measurable_sum _ (λ x _, _), + refine (measurable.ennreal_to_real _).strongly_measurable.smul_const _, + simp only [simple_func.coe_comp, preimage_comp] {single_pass := tt}, + apply measurable_measure_prod_mk_left, + exact (s n).measurable_set_fiber x }, + have h2f' : tendsto f' at_top (𝓝 (λ (x : α), ∫ (y : β), f x y ∂ν)), + { rw [tendsto_pi_nhds], intro x, + by_cases hfx : integrable (f x) ν, + { have : ∀ n, integrable (s' n x) ν, + { intro n, apply (hfx.norm.add hfx.norm).mono' (s' n x).ae_strongly_measurable, + apply eventually_of_forall, intro y, + simp_rw [s', simple_func.coe_comp], exact simple_func.norm_approx_on_zero_le _ _ (x, y) n }, + simp only [f', hfx, simple_func.integral_eq_integral _ (this _), indicator_of_mem, + mem_set_of_eq], + refine tendsto_integral_of_dominated_convergence (λ y, ‖f x y‖ + ‖f x y‖) + (λ n, (s' n x).ae_strongly_measurable) (hfx.norm.add hfx.norm) _ _, + { exact λ n, eventually_of_forall (λ y, simple_func.norm_approx_on_zero_le _ _ (x, y) n) }, + { refine eventually_of_forall (λ y, simple_func.tendsto_approx_on _ _ _), + apply subset_closure, + simp [-uncurry_apply_pair], } }, + { simp [f', hfx, integral_undef], } }, + exact strongly_measurable_of_tendsto _ hf' h2f' +end + +/-- The Bochner integral is measurable. This shows that the integrand of (the right-hand-side of) + Fubini's theorem is measurable. -/ +lemma measure_theory.strongly_measurable.integral_prod_right' [sigma_finite ν] ⦃f : α × β → E⦄ + (hf : strongly_measurable f) : strongly_measurable (λ x, ∫ y, f (x, y) ∂ν) := +by { rw [← uncurry_curry f] at hf, exact hf.integral_prod_right } + +/-- The Bochner integral is measurable. This shows that the integrand of (the right-hand-side of) + the symmetric version of Fubini's theorem is measurable. + This version has `f` in curried form. -/ +lemma measure_theory.strongly_measurable.integral_prod_left [sigma_finite μ] ⦃f : α → β → E⦄ + (hf : strongly_measurable (uncurry f)) : strongly_measurable (λ y, ∫ x, f x y ∂μ) := +(hf.comp_measurable measurable_swap).integral_prod_right' + +/-- The Bochner integral is measurable. This shows that the integrand of (the right-hand-side of) + the symmetric version of Fubini's theorem is measurable. -/ +lemma measure_theory.strongly_measurable.integral_prod_left' [sigma_finite μ] ⦃f : α × β → E⦄ + (hf : strongly_measurable f) : strongly_measurable (λ y, ∫ x, f (x, y) ∂μ) := +(hf.comp_measurable measurable_swap).integral_prod_right' + +end + +/-! ### The product measure -/ + +namespace measure_theory + +namespace measure + +variables [sigma_finite ν] + +lemma integrable_measure_prod_mk_left {s : set (α × β)} + (hs : measurable_set s) (h2s : (μ.prod ν) s ≠ ∞) : + integrable (λ x, (ν (prod.mk x ⁻¹' s)).to_real) μ := +begin + refine ⟨(measurable_measure_prod_mk_left hs).ennreal_to_real.ae_measurable.ae_strongly_measurable, + _⟩, + simp_rw [has_finite_integral, ennnorm_eq_of_real to_real_nonneg], + convert h2s.lt_top using 1, simp_rw [prod_apply hs], apply lintegral_congr_ae, + refine (ae_measure_lt_top hs h2s).mp _, apply eventually_of_forall, intros x hx, + rw [lt_top_iff_ne_top] at hx, simp [of_real_to_real, hx], +end + +end measure + +open measure + +end measure_theory + +open measure_theory.measure + +section + +lemma measure_theory.ae_strongly_measurable.prod_swap + {γ : Type*} [topological_space γ] [sigma_finite μ] [sigma_finite ν] {f : β × α → γ} + (hf : ae_strongly_measurable f (ν.prod μ)) : + ae_strongly_measurable (λ (z : α × β), f z.swap) (μ.prod ν) := +by { rw ← prod_swap at hf, exact hf.comp_measurable measurable_swap } + +lemma measure_theory.ae_strongly_measurable.fst {γ} [topological_space γ] [sigma_finite ν] + {f : α → γ} (hf : ae_strongly_measurable f μ) : + ae_strongly_measurable (λ (z : α × β), f z.1) (μ.prod ν) := +hf.comp_quasi_measure_preserving quasi_measure_preserving_fst + +lemma measure_theory.ae_strongly_measurable.snd {γ} [topological_space γ] [sigma_finite ν] + {f : β → γ} (hf : ae_strongly_measurable f ν) : + ae_strongly_measurable (λ (z : α × β), f z.2) (μ.prod ν) := +hf.comp_quasi_measure_preserving quasi_measure_preserving_snd + +/-- The Bochner integral is a.e.-measurable. + This shows that the integrand of (the right-hand-side of) Fubini's theorem is a.e.-measurable. -/ +lemma measure_theory.ae_strongly_measurable.integral_prod_right' [sigma_finite ν] + [normed_space ℝ E] [complete_space E] + ⦃f : α × β → E⦄ (hf : ae_strongly_measurable f (μ.prod ν)) : + ae_strongly_measurable (λ x, ∫ y, f (x, y) ∂ν) μ := +⟨λ x, ∫ y, hf.mk f (x, y) ∂ν, hf.strongly_measurable_mk.integral_prod_right', + by { filter_upwards [ae_ae_of_ae_prod hf.ae_eq_mk] with _ hx using integral_congr_ae hx }⟩ + +lemma measure_theory.ae_strongly_measurable.prod_mk_left + {γ : Type*} [sigma_finite ν] [topological_space γ] {f : α × β → γ} + (hf : ae_strongly_measurable f (μ.prod ν)) : ∀ᵐ x ∂μ, ae_strongly_measurable (λ y, f (x, y)) ν := +begin + filter_upwards [ae_ae_of_ae_prod hf.ae_eq_mk] with x hx, + exact ⟨λ y, hf.mk f (x, y), hf.strongly_measurable_mk.comp_measurable measurable_prod_mk_left, hx⟩ +end + +end + +namespace measure_theory + +variables [sigma_finite ν] + +/-! ### Integrability on a product -/ +section + +lemma integrable.swap [sigma_finite μ] ⦃f : α × β → E⦄ + (hf : integrable f (μ.prod ν)) : integrable (f ∘ prod.swap) (ν.prod μ) := +⟨hf.ae_strongly_measurable.prod_swap, + (lintegral_prod_swap _ hf.ae_strongly_measurable.ennnorm : _).le.trans_lt hf.has_finite_integral⟩ + +lemma integrable_swap_iff [sigma_finite μ] ⦃f : α × β → E⦄ : + integrable (f ∘ prod.swap) (ν.prod μ) ↔ integrable f (μ.prod ν) := +⟨λ hf, by { convert hf.swap, ext ⟨x, y⟩, refl }, λ hf, hf.swap⟩ + +lemma has_finite_integral_prod_iff ⦃f : α × β → E⦄ (h1f : strongly_measurable f) : + has_finite_integral f (μ.prod ν) ↔ (∀ᵐ x ∂ μ, has_finite_integral (λ y, f (x, y)) ν) ∧ + has_finite_integral (λ x, ∫ y, ‖f (x, y)‖ ∂ν) μ := +begin + simp only [has_finite_integral, lintegral_prod_of_measurable _ h1f.ennnorm], + have : ∀ x, ∀ᵐ y ∂ν, 0 ≤ ‖f (x, y)‖ := λ x, eventually_of_forall (λ y, norm_nonneg _), + simp_rw [integral_eq_lintegral_of_nonneg_ae (this _) + (h1f.norm.comp_measurable measurable_prod_mk_left).ae_strongly_measurable, + ennnorm_eq_of_real to_real_nonneg, of_real_norm_eq_coe_nnnorm], + -- this fact is probably too specialized to be its own lemma + have : ∀ {p q r : Prop} (h1 : r → p), (r ↔ p ∧ q) ↔ (p → (r ↔ q)) := + λ p q r h1, by rw [← and.congr_right_iff, and_iff_right_of_imp h1], + rw [this], + { intro h2f, rw lintegral_congr_ae, + refine h2f.mp _, apply eventually_of_forall, intros x hx, dsimp only, + rw [of_real_to_real], rw [← lt_top_iff_ne_top], exact hx }, + { intro h2f, refine ae_lt_top _ h2f.ne, exact h1f.ennnorm.lintegral_prod_right' }, +end + +lemma has_finite_integral_prod_iff' ⦃f : α × β → E⦄ (h1f : ae_strongly_measurable f (μ.prod ν)) : + has_finite_integral f (μ.prod ν) ↔ (∀ᵐ x ∂ μ, has_finite_integral (λ y, f (x, y)) ν) ∧ + has_finite_integral (λ x, ∫ y, ‖f (x, y)‖ ∂ν) μ := +begin + rw [has_finite_integral_congr h1f.ae_eq_mk, + has_finite_integral_prod_iff h1f.strongly_measurable_mk], + apply and_congr, + { apply eventually_congr, + filter_upwards [ae_ae_of_ae_prod h1f.ae_eq_mk.symm], + assume x hx, + exact has_finite_integral_congr hx }, + { apply has_finite_integral_congr, + filter_upwards [ae_ae_of_ae_prod h1f.ae_eq_mk.symm] with _ hx + using integral_congr_ae (eventually_eq.fun_comp hx _), }, + { apply_instance, }, +end + +/-- A binary function is integrable if the function `y ↦ f (x, y)` is integrable for almost every + `x` and the function `x ↦ ∫ ‖f (x, y)‖ dy` is integrable. -/ +lemma integrable_prod_iff ⦃f : α × β → E⦄ (h1f : ae_strongly_measurable f (μ.prod ν)) : + integrable f (μ.prod ν) ↔ + (∀ᵐ x ∂ μ, integrable (λ y, f (x, y)) ν) ∧ integrable (λ x, ∫ y, ‖f (x, y)‖ ∂ν) μ := +by simp [integrable, h1f, has_finite_integral_prod_iff', h1f.norm.integral_prod_right', + h1f.prod_mk_left] + +/-- A binary function is integrable if the function `x ↦ f (x, y)` is integrable for almost every + `y` and the function `y ↦ ∫ ‖f (x, y)‖ dx` is integrable. -/ +lemma integrable_prod_iff' [sigma_finite μ] ⦃f : α × β → E⦄ + (h1f : ae_strongly_measurable f (μ.prod ν)) : + integrable f (μ.prod ν) ↔ + (∀ᵐ y ∂ ν, integrable (λ x, f (x, y)) μ) ∧ integrable (λ y, ∫ x, ‖f (x, y)‖ ∂μ) ν := +by { convert integrable_prod_iff (h1f.prod_swap) using 1, rw [integrable_swap_iff] } + +lemma integrable.prod_left_ae [sigma_finite μ] ⦃f : α × β → E⦄ + (hf : integrable f (μ.prod ν)) : ∀ᵐ y ∂ ν, integrable (λ x, f (x, y)) μ := +((integrable_prod_iff' hf.ae_strongly_measurable).mp hf).1 + +lemma integrable.prod_right_ae [sigma_finite μ] ⦃f : α × β → E⦄ + (hf : integrable f (μ.prod ν)) : ∀ᵐ x ∂ μ, integrable (λ y, f (x, y)) ν := +hf.swap.prod_left_ae + +lemma integrable.integral_norm_prod_left ⦃f : α × β → E⦄ + (hf : integrable f (μ.prod ν)) : integrable (λ x, ∫ y, ‖f (x, y)‖ ∂ν) μ := +((integrable_prod_iff hf.ae_strongly_measurable).mp hf).2 + +lemma integrable.integral_norm_prod_right [sigma_finite μ] ⦃f : α × β → E⦄ + (hf : integrable f (μ.prod ν)) : integrable (λ y, ∫ x, ‖f (x, y)‖ ∂μ) ν := +hf.swap.integral_norm_prod_left + +lemma integrable_prod_mul {L : Type*} [is_R_or_C L] + {f : α → L} {g : β → L} (hf : integrable f μ) (hg : integrable g ν) : + integrable (λ (z : α × β), f z.1 * g z.2) (μ.prod ν) := +begin + refine (integrable_prod_iff _).2 ⟨_, _⟩, + { exact hf.1.fst.mul hg.1.snd }, + { exact eventually_of_forall (λ x, hg.const_mul (f x)) }, + { simpa only [norm_mul, integral_mul_left] using hf.norm.mul_const _ } +end + +end + +variables [normed_space ℝ E] [complete_space E] + +lemma integrable.integral_prod_left ⦃f : α × β → E⦄ + (hf : integrable f (μ.prod ν)) : integrable (λ x, ∫ y, f (x, y) ∂ν) μ := +integrable.mono hf.integral_norm_prod_left hf.ae_strongly_measurable.integral_prod_right' $ + eventually_of_forall $ λ x, (norm_integral_le_integral_norm _).trans_eq $ + (norm_of_nonneg $ integral_nonneg_of_ae $ eventually_of_forall $ + λ y, (norm_nonneg (f (x, y)) : _)).symm + +lemma integrable.integral_prod_right [sigma_finite μ] ⦃f : α × β → E⦄ + (hf : integrable f (μ.prod ν)) : integrable (λ y, ∫ x, f (x, y) ∂μ) ν := +hf.swap.integral_prod_left + +/-! ### The Bochner integral on a product -/ + +variables [sigma_finite μ] + +lemma integral_prod_swap (f : α × β → E) + (hf : ae_strongly_measurable f (μ.prod ν)) : ∫ z, f z.swap ∂(ν.prod μ) = ∫ z, f z ∂(μ.prod ν) := +begin + rw ← prod_swap at hf, + rw [← integral_map measurable_swap.ae_measurable hf, prod_swap] +end + +variables {E' : Type*} [normed_add_comm_group E'] [complete_space E'] [normed_space ℝ E'] + +/-! Some rules about the sum/difference of double integrals. They follow from `integral_add`, but + we separate them out as separate lemmas, because they involve quite some steps. -/ + +/-- Integrals commute with addition inside another integral. `F` can be any function. -/ +lemma integral_fn_integral_add ⦃f g : α × β → E⦄ (F : E → E') + (hf : integrable f (μ.prod ν)) (hg : integrable g (μ.prod ν)) : + ∫ x, F (∫ y, f (x, y) + g (x, y) ∂ν) ∂μ = ∫ x, F (∫ y, f (x, y) ∂ν + ∫ y, g (x, y) ∂ν) ∂μ := +begin + refine integral_congr_ae _, + filter_upwards [hf.prod_right_ae, hg.prod_right_ae] with _ h2f h2g, + simp [integral_add h2f h2g], +end + +/-- Integrals commute with subtraction inside another integral. + `F` can be any measurable function. -/ +lemma integral_fn_integral_sub ⦃f g : α × β → E⦄ (F : E → E') + (hf : integrable f (μ.prod ν)) (hg : integrable g (μ.prod ν)) : + ∫ x, F (∫ y, f (x, y) - g (x, y) ∂ν) ∂μ = ∫ x, F (∫ y, f (x, y) ∂ν - ∫ y, g (x, y) ∂ν) ∂μ := +begin + refine integral_congr_ae _, + filter_upwards [hf.prod_right_ae, hg.prod_right_ae] with _ h2f h2g, + simp [integral_sub h2f h2g], +end + +/-- Integrals commute with subtraction inside a lower Lebesgue integral. + `F` can be any function. -/ +lemma lintegral_fn_integral_sub ⦃f g : α × β → E⦄ + (F : E → ℝ≥0∞) (hf : integrable f (μ.prod ν)) (hg : integrable g (μ.prod ν)) : + ∫⁻ x, F (∫ y, f (x, y) - g (x, y) ∂ν) ∂μ = ∫⁻ x, F (∫ y, f (x, y) ∂ν - ∫ y, g (x, y) ∂ν) ∂μ := +begin + refine lintegral_congr_ae _, + filter_upwards [hf.prod_right_ae, hg.prod_right_ae] with _ h2f h2g, + simp [integral_sub h2f h2g], +end + +/-- Double integrals commute with addition. -/ +lemma integral_integral_add ⦃f g : α × β → E⦄ + (hf : integrable f (μ.prod ν)) (hg : integrable g (μ.prod ν)) : + ∫ x, ∫ y, f (x, y) + g (x, y) ∂ν ∂μ = ∫ x, ∫ y, f (x, y) ∂ν ∂μ + ∫ x, ∫ y, g (x, y) ∂ν ∂μ := +(integral_fn_integral_add id hf hg).trans $ + integral_add hf.integral_prod_left hg.integral_prod_left + +/-- Double integrals commute with addition. This is the version with `(f + g) (x, y)` + (instead of `f (x, y) + g (x, y)`) in the LHS. -/ +lemma integral_integral_add' ⦃f g : α × β → E⦄ + (hf : integrable f (μ.prod ν)) (hg : integrable g (μ.prod ν)) : + ∫ x, ∫ y, (f + g) (x, y) ∂ν ∂μ = ∫ x, ∫ y, f (x, y) ∂ν ∂μ + ∫ x, ∫ y, g (x, y) ∂ν ∂μ := +integral_integral_add hf hg + +/-- Double integrals commute with subtraction. -/ +lemma integral_integral_sub ⦃f g : α × β → E⦄ + (hf : integrable f (μ.prod ν)) (hg : integrable g (μ.prod ν)) : + ∫ x, ∫ y, f (x, y) - g (x, y) ∂ν ∂μ = ∫ x, ∫ y, f (x, y) ∂ν ∂μ - ∫ x, ∫ y, g (x, y) ∂ν ∂μ := +(integral_fn_integral_sub id hf hg).trans $ + integral_sub hf.integral_prod_left hg.integral_prod_left + +/-- Double integrals commute with subtraction. This is the version with `(f - g) (x, y)` + (instead of `f (x, y) - g (x, y)`) in the LHS. -/ +lemma integral_integral_sub' ⦃f g : α × β → E⦄ + (hf : integrable f (μ.prod ν)) (hg : integrable g (μ.prod ν)) : + ∫ x, ∫ y, (f - g) (x, y) ∂ν ∂μ = ∫ x, ∫ y, f (x, y) ∂ν ∂μ - ∫ x, ∫ y, g (x, y) ∂ν ∂μ := +integral_integral_sub hf hg + +/-- The map that sends an L¹-function `f : α × β → E` to `∫∫f` is continuous. -/ +lemma continuous_integral_integral : + continuous (λ (f : α × β →₁[μ.prod ν] E), ∫ x, ∫ y, f (x, y) ∂ν ∂μ) := +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, (‖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, ‖i (x, y) - g (x, y)‖₊ ∂ν ∂μ }, + swap, { exact λ i, lintegral_mono (λ x, ennnorm_integral_le_lintegral_ennnorm _) }, + show tendsto (λ (i : α × β →₁[μ.prod ν] E), + ∫⁻ 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], + refine (continuous_of_real.tendsto 0).comp _, + rw [← tendsto_iff_norm_tendsto_zero], exact tendsto_id +end + +/-- **Fubini's Theorem**: For integrable functions on `α × β`, + the Bochner integral of `f` is equal to the iterated Bochner integral. + `integrable_prod_iff` can be useful to show that the function in question in integrable. + `measure_theory.integrable.integral_prod_right` is useful to show that the inner integral + of the right-hand side is integrable. -/ +lemma integral_prod : ∀ (f : α × β → E) (hf : integrable f (μ.prod ν)), + ∫ z, f z ∂(μ.prod ν) = ∫ x, ∫ y, f (x, y) ∂ν ∂μ := +begin + apply integrable.induction, + { intros c s hs h2s, + simp_rw [integral_indicator hs, ← indicator_comp_right, + function.comp, integral_indicator (measurable_prod_mk_left hs), + set_integral_const, integral_smul_const, + integral_to_real (measurable_measure_prod_mk_left hs).ae_measurable + (ae_measure_lt_top hs h2s.ne), prod_apply hs] }, + { intros f g hfg i_f i_g hf hg, + simp_rw [integral_add' i_f i_g, integral_integral_add' i_f i_g, hf, hg] }, + { exact is_closed_eq continuous_integral continuous_integral_integral }, + { intros f g hfg i_f hf, convert hf using 1, + { exact integral_congr_ae hfg.symm }, + { refine integral_congr_ae _, + refine (ae_ae_of_ae_prod hfg).mp _, + apply eventually_of_forall, intros x hfgx, + exact integral_congr_ae (ae_eq_symm hfgx) } } +end + +/-- Symmetric version of **Fubini's Theorem**: For integrable functions on `α × β`, + the Bochner integral of `f` is equal to the iterated Bochner integral. + This version has the integrals on the right-hand side in the other order. -/ +lemma integral_prod_symm (f : α × β → E) (hf : integrable f (μ.prod ν)) : + ∫ z, f z ∂(μ.prod ν) = ∫ y, ∫ x, f (x, y) ∂μ ∂ν := +by { simp_rw [← integral_prod_swap f hf.ae_strongly_measurable], exact integral_prod _ hf.swap } + +/-- Reversed version of **Fubini's Theorem**. -/ +lemma integral_integral {f : α → β → E} (hf : integrable (uncurry f) (μ.prod ν)) : + ∫ x, ∫ y, f x y ∂ν ∂μ = ∫ z, f z.1 z.2 ∂(μ.prod ν) := +(integral_prod _ hf).symm + +/-- Reversed version of **Fubini's Theorem** (symmetric version). -/ +lemma integral_integral_symm {f : α → β → E} (hf : integrable (uncurry f) (μ.prod ν)) : + ∫ x, ∫ y, f x y ∂ν ∂μ = ∫ z, f z.2 z.1 ∂(ν.prod μ) := +(integral_prod_symm _ hf.swap).symm + +/-- Change the order of Bochner integration. -/ +lemma integral_integral_swap ⦃f : α → β → E⦄ (hf : integrable (uncurry f) (μ.prod ν)) : + ∫ x, ∫ y, f x y ∂ν ∂μ = ∫ y, ∫ x, f x y ∂μ ∂ν := +(integral_integral hf).trans (integral_prod_symm _ hf) + +/-- **Fubini's Theorem** for set integrals. -/ +lemma set_integral_prod (f : α × β → E) {s : set α} {t : set β} + (hf : integrable_on f (s ×ˢ t) (μ.prod ν)) : + ∫ z in s ×ˢ t, f z ∂(μ.prod ν) = ∫ x in s, ∫ y in t, f (x, y) ∂ν ∂μ := +begin + simp only [← measure.prod_restrict s t, integrable_on] at hf ⊢, + exact integral_prod f hf +end + +lemma integral_prod_mul {L : Type*} [is_R_or_C L] (f : α → L) (g : β → L) : + ∫ z, f z.1 * g z.2 ∂(μ.prod ν) = (∫ x, f x ∂μ) * (∫ y, g y ∂ν) := +begin + by_cases h : integrable (λ (z : α × β), f z.1 * g z.2) (μ.prod ν), + { rw integral_prod _ h, + simp_rw [integral_mul_left, integral_mul_right] }, + have H : ¬(integrable f μ) ∨ ¬(integrable g ν), + { contrapose! h, + exact integrable_prod_mul h.1 h.2 }, + cases H; + simp [integral_undef h, integral_undef H], +end + +lemma set_integral_prod_mul {L : Type*} [is_R_or_C L] + (f : α → L) (g : β → L) (s : set α) (t : set β) : + ∫ z in s ×ˢ t, f z.1 * g z.2 ∂(μ.prod ν) = (∫ x in s, f x ∂μ) * (∫ y in t, g y ∂ν) := +by simp only [← measure.prod_restrict s t, integrable_on, integral_prod_mul] + +end measure_theory diff --git a/src/measure_theory/covering/besicovitch_vector_space.lean b/src/measure_theory/covering/besicovitch_vector_space.lean index 9cba2ebdeafdb..fab628357d0ab 100644 --- a/src/measure_theory/covering/besicovitch_vector_space.lean +++ b/src/measure_theory/covering/besicovitch_vector_space.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ -import measure_theory.measure.haar_lebesgue +import measure_theory.measure.lebesgue.eq_haar import measure_theory.covering.besicovitch /-! diff --git a/src/measure_theory/covering/one_dim.lean b/src/measure_theory/covering/one_dim.lean index 1df64bd0d094b..305f17edf8f87 100644 --- a/src/measure_theory/covering/one_dim.lean +++ b/src/measure_theory/covering/one_dim.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ import measure_theory.covering.density_theorem -import measure_theory.measure.haar_lebesgue +import measure_theory.measure.lebesgue.eq_haar /-! # Covering theorems for Lebesgue measure in one dimension diff --git a/src/measure_theory/function/jacobian.lean b/src/measure_theory/function/jacobian.lean index c90328ecd81bd..52a765f701dd1 100644 --- a/src/measure_theory/function/jacobian.lean +++ b/src/measure_theory/function/jacobian.lean @@ -6,7 +6,7 @@ Authors: Sébastien Gouëzel import analysis.calculus.inverse import measure_theory.constructions.borel_space.continuous_linear_map import measure_theory.covering.besicovitch_vector_space -import measure_theory.measure.haar_lebesgue +import measure_theory.measure.lebesgue.eq_haar import analysis.normed_space.pointwise import measure_theory.constructions.polish diff --git a/src/measure_theory/group/geometry_of_numbers.lean b/src/measure_theory/group/geometry_of_numbers.lean index 0af202df3f364..57d4ce1720ec1 100644 --- a/src/measure_theory/group/geometry_of_numbers.lean +++ b/src/measure_theory/group/geometry_of_numbers.lean @@ -5,7 +5,7 @@ Authors: Alex J. Best -/ import analysis.convex.measure import measure_theory.group.fundamental_domain -import measure_theory.measure.haar_lebesgue +import measure_theory.measure.lebesgue.eq_haar /-! # Geometry of numbers diff --git a/src/measure_theory/group/measure.lean b/src/measure_theory/group/measure.lean index a982613517d5b..8ed1592e22f69 100644 --- a/src/measure_theory/group/measure.lean +++ b/src/measure_theory/group/measure.lean @@ -8,7 +8,7 @@ import measure_theory.measure.regular import measure_theory.group.measurable_equiv import measure_theory.measure.open_pos import measure_theory.group.action -import measure_theory.constructions.prod +import measure_theory.constructions.prod.basic import topology.continuous_function.cocompact_map /-! @@ -727,30 +727,8 @@ begin exact ge_of_tendsto' J I, end -/- The above instance applies in particular to show that an additive Haar measure on a nontrivial -finite-dimensional real vector space has no atom. -/ -example {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] [nontrivial E] - [finite_dimensional ℝ E] [measurable_space E] [borel_space E] (μ : measure E) - [is_add_haar_measure μ] : - has_no_atoms μ := by apply_instance - end -variables [nontrivially_normed_field 𝕜] [topological_space G] [topological_space H] - [add_comm_group G] [add_comm_group H] [topological_add_group G] [topological_add_group H] - [module 𝕜 G] [module 𝕜 H] (μ : measure G) [is_add_haar_measure μ] [borel_space G] [borel_space H] - [t2_space H] - -instance map_continuous_linear_equiv.is_add_haar_measure (e : G ≃L[𝕜] H) : - is_add_haar_measure (μ.map e) := -e.to_add_equiv.is_add_haar_measure_map _ e.continuous e.symm.continuous - -variables [complete_space 𝕜] [t2_space G] [finite_dimensional 𝕜 G] [has_continuous_smul 𝕜 G] - [has_continuous_smul 𝕜 H] - -instance map_linear_equiv.is_add_haar_measure (e : G ≃ₗ[𝕜] H) : is_add_haar_measure (μ.map e) := -map_continuous_linear_equiv.is_add_haar_measure _ e.to_continuous_linear_equiv - end measure end haar diff --git a/src/measure_theory/group/prod.lean b/src/measure_theory/group/prod.lean index c2ad58d818513..20da19bf2aed6 100644 --- a/src/measure_theory/group/prod.lean +++ b/src/measure_theory/group/prod.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Floris van Doorn. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn -/ -import measure_theory.constructions.prod +import measure_theory.constructions.prod.basic import measure_theory.group.measure /-! diff --git a/src/measure_theory/integral/divergence_theorem.lean b/src/measure_theory/integral/divergence_theorem.lean index 465737ffa7277..f0b9fa18f81d5 100644 --- a/src/measure_theory/integral/divergence_theorem.lean +++ b/src/measure_theory/integral/divergence_theorem.lean @@ -6,6 +6,7 @@ Authors: Yury Kudryashov import analysis.box_integral.divergence_theorem import analysis.box_integral.integrability import analysis.calculus.deriv +import measure_theory.constructions.prod.integral import measure_theory.integral.interval_integral /-! diff --git a/src/measure_theory/integral/interval_integral.lean b/src/measure_theory/integral/interval_integral.lean index 15062741e82dc..9c446ff60f04a 100644 --- a/src/measure_theory/integral/interval_integral.lean +++ b/src/measure_theory/integral/interval_integral.lean @@ -5,7 +5,7 @@ Authors: Yury G. Kudryashov, Patrick Massot, Sébastien Gouëzel -/ import data.set.intervals.disjoint import measure_theory.integral.set_integral -import measure_theory.measure.lebesgue +import measure_theory.measure.lebesgue.basic /-! # Integral over an interval diff --git a/src/measure_theory/integral/periodic.lean b/src/measure_theory/integral/periodic.lean index 48ed554e4831a..9c87a4e07a0bc 100644 --- a/src/measure_theory/integral/periodic.lean +++ b/src/measure_theory/integral/periodic.lean @@ -3,8 +3,8 @@ Copyright (c) 2021 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov, Alex Kontorovich, Heather Macbeth -/ -import measure_theory.measure.haar_lebesgue -import measure_theory.measure.haar_quotient +import measure_theory.measure.lebesgue.eq_haar +import measure_theory.measure.haar.quotient import measure_theory.integral.interval_integral import topology.algebra.order.floor diff --git a/src/measure_theory/integral/torus_integral.lean b/src/measure_theory/integral/torus_integral.lean index 2d839e67cbd38..02151c688526a 100644 --- a/src/measure_theory/integral/torus_integral.lean +++ b/src/measure_theory/integral/torus_integral.lean @@ -3,6 +3,7 @@ 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.constructions.prod.integral import measure_theory.integral.circle_integral /-! diff --git a/src/measure_theory/measure/haar.lean b/src/measure_theory/measure/haar/basic.lean similarity index 100% rename from src/measure_theory/measure/haar.lean rename to src/measure_theory/measure/haar/basic.lean diff --git a/src/measure_theory/measure/haar_of_inner.lean b/src/measure_theory/measure/haar/inner_product_space.lean similarity index 98% rename from src/measure_theory/measure/haar_of_inner.lean rename to src/measure_theory/measure/haar/inner_product_space.lean index 0923810fae711..999edc737235e 100644 --- a/src/measure_theory/measure/haar_of_inner.lean +++ b/src/measure_theory/measure/haar/inner_product_space.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ import analysis.inner_product_space.orientation -import measure_theory.measure.haar_lebesgue +import measure_theory.measure.lebesgue.eq_haar /-! # Volume forms and measures on inner product spaces diff --git a/src/measure_theory/measure/haar/normed_space.lean b/src/measure_theory/measure/haar/normed_space.lean new file mode 100644 index 0000000000000..5fe6c03c96386 --- /dev/null +++ b/src/measure_theory/measure/haar/normed_space.lean @@ -0,0 +1,102 @@ +/- +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, Sébastien Gouëzel +-/ +import measure_theory.measure.lebesgue.eq_haar +import measure_theory.integral.bochner + +/-! +# Basic properties of Haar measures on real vector spaces + +-/ + +noncomputable theory + +open_locale nnreal ennreal pointwise big_operators topology +open has_inv set function measure_theory.measure filter +open measure finite_dimensional + +namespace measure_theory + +namespace measure + +/- The instance `is_add_haar_measure.has_no_atoms` applies in particular to show that an additive +Haar measure on a nontrivial finite-dimensional real vector space has no atom. -/ +example {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] [nontrivial E] + [finite_dimensional ℝ E] [measurable_space E] [borel_space E] (μ : measure E) + [is_add_haar_measure μ] : + has_no_atoms μ := by apply_instance + +section continuous_linear_equiv + +variables {𝕜 G H : Type*} [measurable_space G] [measurable_space H] + [nontrivially_normed_field 𝕜] [topological_space G] [topological_space H] + [add_comm_group G] [add_comm_group H] [topological_add_group G] [topological_add_group H] + [module 𝕜 G] [module 𝕜 H] (μ : measure G) [is_add_haar_measure μ] [borel_space G] [borel_space H] + [t2_space H] + +instance map_continuous_linear_equiv.is_add_haar_measure (e : G ≃L[𝕜] H) : + is_add_haar_measure (μ.map e) := +e.to_add_equiv.is_add_haar_measure_map _ e.continuous e.symm.continuous + +variables [complete_space 𝕜] [t2_space G] [finite_dimensional 𝕜 G] [has_continuous_smul 𝕜 G] + [has_continuous_smul 𝕜 H] + +instance map_linear_equiv.is_add_haar_measure (e : G ≃ₗ[𝕜] H) : is_add_haar_measure (μ.map e) := +map_continuous_linear_equiv.is_add_haar_measure _ e.to_continuous_linear_equiv + +end continuous_linear_equiv + +variables {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] [measurable_space E] + [borel_space E] [finite_dimensional ℝ E] (μ : measure E) [is_add_haar_measure μ] + {F : Type*} [normed_add_comm_group F] [normed_space ℝ F] [complete_space F] +variables (μ) {s : set E} + +/-- The integral of `f (R • x)` with respect to an additive Haar measure is a multiple of the +integral of `f`. The formula we give works even when `f` is not integrable or `R = 0` +thanks to the convention that a non-integrable function has integral zero. -/ +lemma integral_comp_smul (f : E → F) (R : ℝ) : + ∫ x, f (R • x) ∂μ = |(R ^ finrank ℝ E)⁻¹| • ∫ x, f x ∂μ := +begin + rcases eq_or_ne R 0 with rfl|hR, + { simp only [zero_smul, integral_const], + rcases nat.eq_zero_or_pos (finrank ℝ E) with hE|hE, + { haveI : subsingleton E, from finrank_zero_iff.1 hE, + have : f = (λ x, f 0), { ext x, rw subsingleton.elim x 0 }, + conv_rhs { rw this }, + simp only [hE, pow_zero, inv_one, abs_one, one_smul, integral_const] }, + { haveI : nontrivial E, from finrank_pos_iff.1 hE, + simp only [zero_pow hE, measure_univ_of_is_add_left_invariant, ennreal.top_to_real, zero_smul, + inv_zero, abs_zero]} }, + { calc ∫ x, f (R • x) ∂μ = ∫ y, f y ∂(measure.map (λ x, R • x) μ) : + (integral_map_equiv (homeomorph.smul (is_unit_iff_ne_zero.2 hR).unit) + .to_measurable_equiv f).symm + ... = |(R ^ finrank ℝ E)⁻¹| • ∫ x, f x ∂μ : + by simp only [map_add_haar_smul μ hR, integral_smul_measure, ennreal.to_real_of_real, + abs_nonneg] } +end + +/-- The integral of `f (R • x)` with respect to an additive Haar measure is a multiple of the +integral of `f`. The formula we give works even when `f` is not integrable or `R = 0` +thanks to the convention that a non-integrable function has integral zero. -/ +lemma integral_comp_smul_of_nonneg (f : E → F) (R : ℝ) {hR : 0 ≤ R} : + ∫ x, f (R • x) ∂μ = (R ^ finrank ℝ E)⁻¹ • ∫ x, f x ∂μ := +by rw [integral_comp_smul μ f R, abs_of_nonneg (inv_nonneg.2 (pow_nonneg hR _))] + +/-- The integral of `f (R⁻¹ • x)` with respect to an additive Haar measure is a multiple of the +integral of `f`. The formula we give works even when `f` is not integrable or `R = 0` +thanks to the convention that a non-integrable function has integral zero. -/ +lemma integral_comp_inv_smul (f : E → F) (R : ℝ) : + ∫ x, f (R⁻¹ • x) ∂μ = |(R ^ finrank ℝ E)| • ∫ x, f x ∂μ := +by rw [integral_comp_smul μ f (R⁻¹), inv_pow, inv_inv] + +/-- The integral of `f (R⁻¹ • x)` with respect to an additive Haar measure is a multiple of the +integral of `f`. The formula we give works even when `f` is not integrable or `R = 0` +thanks to the convention that a non-integrable function has integral zero. -/ +lemma integral_comp_inv_smul_of_nonneg (f : E → F) {R : ℝ} (hR : 0 ≤ R) : + ∫ x, f (R⁻¹ • x) ∂μ = R ^ finrank ℝ E • ∫ x, f x ∂μ := +by rw [integral_comp_inv_smul μ f R, abs_of_nonneg ((pow_nonneg hR _))] + +end measure +end measure_theory diff --git a/src/measure_theory/measure/haar_of_basis.lean b/src/measure_theory/measure/haar/of_basis.lean similarity index 99% rename from src/measure_theory/measure/haar_of_basis.lean rename to src/measure_theory/measure/haar/of_basis.lean index 3da1e80196ca2..0c0cf1c9bb18f 100644 --- a/src/measure_theory/measure/haar_of_basis.lean +++ b/src/measure_theory/measure/haar/of_basis.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ -import measure_theory.measure.haar +import measure_theory.measure.haar.basic import analysis.inner_product_space.pi_L2 /-! diff --git a/src/measure_theory/measure/haar_quotient.lean b/src/measure_theory/measure/haar/quotient.lean similarity index 99% rename from src/measure_theory/measure/haar_quotient.lean rename to src/measure_theory/measure/haar/quotient.lean index 67370553b3eda..87116a3ea240e 100644 --- a/src/measure_theory/measure/haar_quotient.lean +++ b/src/measure_theory/measure/haar/quotient.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Alex Kontorovich, Heather Macbeth -/ -import measure_theory.measure.haar +import measure_theory.measure.haar.basic import measure_theory.group.fundamental_domain import algebra.group.opposite diff --git a/src/measure_theory/measure/hausdorff.lean b/src/measure_theory/measure/hausdorff.lean index aa1cf0a4df7da..815aab7f5eb93 100644 --- a/src/measure_theory/measure/hausdorff.lean +++ b/src/measure_theory/measure/hausdorff.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import measure_theory.constructions.borel_space.basic -import measure_theory.measure.lebesgue +import measure_theory.measure.lebesgue.basic import topology.metric_space.holder import topology.metric_space.metric_separated diff --git a/src/measure_theory/measure/lebesgue.lean b/src/measure_theory/measure/lebesgue/basic.lean similarity index 88% rename from src/measure_theory/measure/lebesgue.lean rename to src/measure_theory/measure/lebesgue/basic.lean index c347ccce0c02b..01c2ff63e5f03 100644 --- a/src/measure_theory/measure/lebesgue.lean +++ b/src/measure_theory/measure/lebesgue/basic.lean @@ -9,7 +9,7 @@ import linear_algebra.matrix.diagonal import linear_algebra.matrix.transvection import measure_theory.constructions.pi import measure_theory.measure.stieltjes -import measure_theory.measure.haar_of_basis +import measure_theory.measure.haar.of_basis /-! # Lebesgue measure on the real line and on `ℝⁿ` @@ -22,10 +22,12 @@ Lebesgue measure on `ℝⁿ`. In particular, we prove that they are translation We show that, on `ℝⁿ`, a linear map acts on Lebesgue measure by rescaling it through the absolute value of its determinant, in `real.map_linear_map_volume_pi_eq_smul_volume_pi`. -More properties of the Lebesgue measure are deduced from this in `haar_lebesgue.lean`, where they +More properties of the Lebesgue measure are deduced from this in `lebesgue.eq_haar.lean`, where they are proved more generally for any additive Haar measure on a finite-dimensional real vector space. -/ +assert_not_exists measure_theory.integral + noncomputable theory open classical set filter measure_theory measure_theory.measure topological_space open ennreal (of_real) @@ -505,29 +507,6 @@ begin (region_between_subset (ae_measurable.mk f hf) (ae_measurable.mk g hg) s)).symm }, end -theorem volume_region_between_eq_integral' [sigma_finite μ] - (f_int : integrable_on f s μ) (g_int : integrable_on g s μ) - (hs : measurable_set s) (hfg : f ≤ᵐ[μ.restrict s] g ) : - μ.prod volume (region_between f g s) = ennreal.of_real (∫ y in s, (g - f) y ∂μ) := -begin - have h : g - f =ᵐ[μ.restrict s] (λ x, real.to_nnreal (g x - f x)), - from hfg.mono (λ x hx, (real.coe_to_nnreal _ $ sub_nonneg.2 hx).symm), - rw [volume_region_between_eq_lintegral f_int.ae_measurable g_int.ae_measurable hs, - integral_congr_ae h, lintegral_congr_ae, - lintegral_coe_eq_integral _ ((integrable_congr h).mp (g_int.sub f_int))], - simpa only, -end - -/-- If two functions are integrable on a measurable set, and one function is less than - or equal to the other on that set, then the volume of the region - between the two functions can be represented as an integral. -/ -theorem volume_region_between_eq_integral [sigma_finite μ] - (f_int : integrable_on f s μ) (g_int : integrable_on g s μ) - (hs : measurable_set s) (hfg : ∀ x ∈ s, f x ≤ g x) : - μ.prod volume (region_between f g s) = ennreal.of_real (∫ y in s, (g - f) y ∂μ) := -volume_region_between_eq_integral' f_int g_int hs - ((ae_restrict_iff' hs).mpr (eventually_of_forall hfg)) - end region_between /-- Consider a real set `s`. If a property is true almost everywhere in `s ∩ (a, b)` for @@ -601,59 +580,3 @@ begin apply h'x p pA ⟨xs, xp⟩ }, { exact false.elim (hx ⟨xs, Hx⟩) } end - -section summable_norm_Icc - -open continuous_map - -/- The following lemma is a minor variation on `integrable_of_summable_norm_restrict` in -`measure_theory.integral.set_integral`, but it is placed here because it needs to know that -`Icc a b` has volume `b - a`. -/ - -/-- If the sequence with `n`-th term the the sup norm of `λ x, f (x + n)` on the interval `Icc 0 1`, -for `n ∈ ℤ`, is summable, then `f` is integrable on `ℝ`. -/ -lemma real.integrable_of_summable_norm_Icc {E : Type*} [normed_add_comm_group E] {f : C(ℝ, E)} - (hf : summable (λ n : ℤ, ‖(f.comp $ continuous_map.add_right n).restrict (Icc 0 1)‖)) : - integrable f := -begin - refine integrable_of_summable_norm_restrict (summable_of_nonneg_of_le - (λ n : ℤ, mul_nonneg (norm_nonneg (f.restrict (⟨Icc n (n + 1), is_compact_Icc⟩ : compacts ℝ))) - ennreal.to_real_nonneg) (λ n, _) hf) (Union_Icc_int_cast ℝ), - simp only [compacts.coe_mk, real.volume_Icc, add_sub_cancel', ennreal.to_real_of_real zero_le_one, - mul_one, norm_le _ (norm_nonneg _)], - intro x, - have := ((f.comp $ continuous_map.add_right n).restrict (Icc 0 1)).norm_coe_le_norm - ⟨x - n, ⟨sub_nonneg.mpr x.2.1, sub_le_iff_le_add'.mpr x.2.2⟩⟩, - simpa only [continuous_map.restrict_apply, comp_apply, coe_add_right, subtype.coe_mk, - sub_add_cancel] - using this, -end - -end summable_norm_Icc - -/-! -### Substituting `-x` for `x` - -These lemmas are stated in terms of either `Iic` or `Ioi` (neglecting `Iio` and `Ici`) to match -mathlib's conventions for integrals over finite intervals (see `interval_integral`). For the case -of finite integrals, see `interval_integral.integral_comp_neg`. --/ - -@[simp] lemma integral_comp_neg_Iic {E : Type*} - [normed_add_comm_group E] [normed_space ℝ E] [complete_space E] (c : ℝ) (f : ℝ → E) : - ∫ x in Iic c, f (-x) = ∫ x in Ioi (-c), f x := -begin - have A : measurable_embedding (λ x : ℝ, -x), - from (homeomorph.neg ℝ).closed_embedding.measurable_embedding, - have := A.set_integral_map f (Ici (-c)), - rw measure.map_neg_eq_self (volume : measure ℝ) at this, - simp_rw [←integral_Ici_eq_integral_Ioi, this, neg_preimage, preimage_neg_Ici, neg_neg], -end - -@[simp] lemma integral_comp_neg_Ioi {E : Type*} - [normed_add_comm_group E] [normed_space ℝ E] [complete_space E] (c : ℝ) (f : ℝ → E) : - ∫ x in Ioi c, f (-x) = ∫ x in Iic (-c), f x := -begin - rw [←neg_neg c, ←integral_comp_neg_Iic], - simp only [neg_neg], -end diff --git a/src/measure_theory/measure/complex_lebesgue.lean b/src/measure_theory/measure/lebesgue/complex.lean similarity index 91% rename from src/measure_theory/measure/complex_lebesgue.lean rename to src/measure_theory/measure/lebesgue/complex.lean index 07066e07eec81..7f201d245816b 100644 --- a/src/measure_theory/measure/complex_lebesgue.lean +++ b/src/measure_theory/measure/lebesgue/complex.lean @@ -3,8 +3,9 @@ Copyright (c) 2021 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ -import measure_theory.measure.lebesgue -import measure_theory.measure.haar_of_basis +import measure_theory.constructions.borel_space.complex +import measure_theory.measure.lebesgue.basic +import measure_theory.measure.haar.of_basis /-! # Lebesgue measure on `ℂ` diff --git a/src/measure_theory/measure/haar_lebesgue.lean b/src/measure_theory/measure/lebesgue/eq_haar.lean similarity index 93% rename from src/measure_theory/measure/haar_lebesgue.lean rename to src/measure_theory/measure/lebesgue/eq_haar.lean index 2e82c6dc5df52..1bdac70427224 100644 --- a/src/measure_theory/measure/haar_lebesgue.lean +++ b/src/measure_theory/measure/lebesgue/eq_haar.lean @@ -3,11 +3,11 @@ Copyright (c) 2021 Floris van Doorn. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn, Sébastien Gouëzel -/ -import measure_theory.measure.lebesgue -import measure_theory.measure.haar -import linear_algebra.finite_dimensional import analysis.normed_space.pointwise +import linear_algebra.finite_dimensional import measure_theory.group.pointwise +import measure_theory.measure.lebesgue.basic +import measure_theory.measure.haar.basic import measure_theory.measure.doubling /-! @@ -41,6 +41,8 @@ density one for the rescaled copies `{x} + r • t` of a given set `t` with posi small `r`, see `eventually_nonempty_inter_smul_of_density_one`. -/ +assert_not_exists measure_theory.integral + open topological_space set filter metric open_locale ennreal pointwise topology nnreal @@ -388,51 +390,6 @@ calc μ (affine_map.homothety x r '' s) = μ ((λ y, y + x) '' (r • ((λ y, y ... = ennreal.of_real (abs (r ^ finrank ℝ E)) * μ s : by simp only [image_add_right, measure_preimage_add_right, add_haar_smul] -/-- The integral of `f (R • x)` with respect to an additive Haar measure is a multiple of the -integral of `f`. The formula we give works even when `f` is not integrable or `R = 0` -thanks to the convention that a non-integrable function has integral zero. -/ -lemma integral_comp_smul (f : E → F) (R : ℝ) : - ∫ x, f (R • x) ∂μ = |(R ^ finrank ℝ E)⁻¹| • ∫ x, f x ∂μ := -begin - rcases eq_or_ne R 0 with rfl|hR, - { simp only [zero_smul, integral_const], - rcases nat.eq_zero_or_pos (finrank ℝ E) with hE|hE, - { haveI : subsingleton E, from finrank_zero_iff.1 hE, - have : f = (λ x, f 0), { ext x, rw subsingleton.elim x 0 }, - conv_rhs { rw this }, - simp only [hE, pow_zero, inv_one, abs_one, one_smul, integral_const] }, - { haveI : nontrivial E, from finrank_pos_iff.1 hE, - simp only [zero_pow hE, measure_univ_of_is_add_left_invariant, ennreal.top_to_real, zero_smul, - inv_zero, abs_zero]} }, - { calc ∫ x, f (R • x) ∂μ = ∫ y, f y ∂(measure.map (λ x, R • x) μ) : - (integral_map_equiv (homeomorph.smul (is_unit_iff_ne_zero.2 hR).unit) - .to_measurable_equiv f).symm - ... = |(R ^ finrank ℝ E)⁻¹| • ∫ x, f x ∂μ : - by simp only [map_add_haar_smul μ hR, integral_smul_measure, ennreal.to_real_of_real, - abs_nonneg] } -end - -/-- The integral of `f (R • x)` with respect to an additive Haar measure is a multiple of the -integral of `f`. The formula we give works even when `f` is not integrable or `R = 0` -thanks to the convention that a non-integrable function has integral zero. -/ -lemma integral_comp_smul_of_nonneg (f : E → F) (R : ℝ) {hR : 0 ≤ R} : - ∫ x, f (R • x) ∂μ = (R ^ finrank ℝ E)⁻¹ • ∫ x, f x ∂μ := -by rw [integral_comp_smul μ f R, abs_of_nonneg (inv_nonneg.2 (pow_nonneg hR _))] - -/-- The integral of `f (R⁻¹ • x)` with respect to an additive Haar measure is a multiple of the -integral of `f`. The formula we give works even when `f` is not integrable or `R = 0` -thanks to the convention that a non-integrable function has integral zero. -/ -lemma integral_comp_inv_smul (f : E → F) (R : ℝ) : - ∫ x, f (R⁻¹ • x) ∂μ = |(R ^ finrank ℝ E)| • ∫ x, f x ∂μ := -by rw [integral_comp_smul μ f (R⁻¹), inv_pow, inv_inv] - -/-- The integral of `f (R⁻¹ • x)` with respect to an additive Haar measure is a multiple of the -integral of `f`. The formula we give works even when `f` is not integrable or `R = 0` -thanks to the convention that a non-integrable function has integral zero. -/ -lemma integral_comp_inv_smul_of_nonneg (f : E → F) {R : ℝ} (hR : 0 ≤ R) : - ∫ x, f (R⁻¹ • x) ∂μ = R ^ finrank ℝ E • ∫ x, f x ∂μ := -by rw [integral_comp_inv_smul μ f R, abs_of_nonneg ((pow_nonneg hR _))] - /-! We don't need to state `map_add_haar_neg` here, because it has already been proved for general Haar measures on general commutative groups. -/ diff --git a/src/measure_theory/measure/lebesgue/integral.lean b/src/measure_theory/measure/lebesgue/integral.lean new file mode 100644 index 0000000000000..a7542af910607 --- /dev/null +++ b/src/measure_theory/measure/lebesgue/integral.lean @@ -0,0 +1,97 @@ +/- +Copyright (c) 2017 Johannes Hölzl. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johannes Hölzl, Sébastien Gouëzel, Yury Kudryashov +-/ +import measure_theory.integral.set_integral +import measure_theory.measure.lebesgue.basic + +/-! # Properties of integration with respect to the Lebesgue measure -/ + +open set filter measure_theory measure_theory.measure topological_space + +section region_between + +variable {α : Type*} +variables [measurable_space α] {μ : measure α} {f g : α → ℝ} {s : set α} + +theorem volume_region_between_eq_integral' [sigma_finite μ] + (f_int : integrable_on f s μ) (g_int : integrable_on g s μ) + (hs : measurable_set s) (hfg : f ≤ᵐ[μ.restrict s] g ) : + μ.prod volume (region_between f g s) = ennreal.of_real (∫ y in s, (g - f) y ∂μ) := +begin + have h : g - f =ᵐ[μ.restrict s] (λ x, real.to_nnreal (g x - f x)), + from hfg.mono (λ x hx, (real.coe_to_nnreal _ $ sub_nonneg.2 hx).symm), + rw [volume_region_between_eq_lintegral f_int.ae_measurable g_int.ae_measurable hs, + integral_congr_ae h, lintegral_congr_ae, + lintegral_coe_eq_integral _ ((integrable_congr h).mp (g_int.sub f_int))], + simpa only, +end + +/-- If two functions are integrable on a measurable set, and one function is less than + or equal to the other on that set, then the volume of the region + between the two functions can be represented as an integral. -/ +theorem volume_region_between_eq_integral [sigma_finite μ] + (f_int : integrable_on f s μ) (g_int : integrable_on g s μ) + (hs : measurable_set s) (hfg : ∀ x ∈ s, f x ≤ g x) : + μ.prod volume (region_between f g s) = ennreal.of_real (∫ y in s, (g - f) y ∂μ) := +volume_region_between_eq_integral' f_int g_int hs + ((ae_restrict_iff' hs).mpr (eventually_of_forall hfg)) + +end region_between + +section summable_norm_Icc + +open continuous_map + +/- The following lemma is a minor variation on `integrable_of_summable_norm_restrict` in +`measure_theory.integral.set_integral`, but it is placed here because it needs to know that +`Icc a b` has volume `b - a`. -/ + +/-- If the sequence with `n`-th term the the sup norm of `λ x, f (x + n)` on the interval `Icc 0 1`, +for `n ∈ ℤ`, is summable, then `f` is integrable on `ℝ`. -/ +lemma real.integrable_of_summable_norm_Icc {E : Type*} [normed_add_comm_group E] {f : C(ℝ, E)} + (hf : summable (λ n : ℤ, ‖(f.comp $ continuous_map.add_right n).restrict (Icc 0 1)‖)) : + integrable f := +begin + refine integrable_of_summable_norm_restrict (summable_of_nonneg_of_le + (λ n : ℤ, mul_nonneg (norm_nonneg (f.restrict (⟨Icc n (n + 1), is_compact_Icc⟩ : compacts ℝ))) + ennreal.to_real_nonneg) (λ n, _) hf) (Union_Icc_int_cast ℝ), + simp only [compacts.coe_mk, real.volume_Icc, add_sub_cancel', ennreal.to_real_of_real zero_le_one, + mul_one, norm_le _ (norm_nonneg _)], + intro x, + have := ((f.comp $ continuous_map.add_right n).restrict (Icc 0 1)).norm_coe_le_norm + ⟨x - n, ⟨sub_nonneg.mpr x.2.1, sub_le_iff_le_add'.mpr x.2.2⟩⟩, + simpa only [continuous_map.restrict_apply, comp_apply, coe_add_right, subtype.coe_mk, + sub_add_cancel] + using this, +end + +end summable_norm_Icc + +/-! +### Substituting `-x` for `x` + +These lemmas are stated in terms of either `Iic` or `Ioi` (neglecting `Iio` and `Ici`) to match +mathlib's conventions for integrals over finite intervals (see `interval_integral`). For the case +of finite integrals, see `interval_integral.integral_comp_neg`. +-/ + +@[simp] lemma integral_comp_neg_Iic {E : Type*} + [normed_add_comm_group E] [normed_space ℝ E] [complete_space E] (c : ℝ) (f : ℝ → E) : + ∫ x in Iic c, f (-x) = ∫ x in Ioi (-c), f x := +begin + have A : measurable_embedding (λ x : ℝ, -x), + from (homeomorph.neg ℝ).closed_embedding.measurable_embedding, + have := A.set_integral_map f (Ici (-c)), + rw measure.map_neg_eq_self (volume : measure ℝ) at this, + simp_rw [←integral_Ici_eq_integral_Ioi, this, neg_preimage, preimage_neg_Ici, neg_neg], +end + +@[simp] lemma integral_comp_neg_Ioi {E : Type*} + [normed_add_comm_group E] [normed_space ℝ E] [complete_space E] (c : ℝ) (f : ℝ → E) : + ∫ x in Ioi c, f (-x) = ∫ x in Iic (-c), f x := +begin + rw [←neg_neg c, ←integral_comp_neg_Iic], + simp only [neg_neg], +end diff --git a/src/measure_theory/measure/portmanteau.lean b/src/measure_theory/measure/portmanteau.lean index debe09934a0d4..24f012aeba70e 100644 --- a/src/measure_theory/measure/portmanteau.lean +++ b/src/measure_theory/measure/portmanteau.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kalle Kytölä -/ import measure_theory.measure.probability_measure -import measure_theory.measure.lebesgue +import measure_theory.measure.lebesgue.basic /-! # Characterizations of weak convergence of finite measures and probability measures diff --git a/src/number_theory/liouville/measure.lean b/src/number_theory/liouville/measure.lean index c7527bdeb75e6..8e5eb260b4453 100644 --- a/src/number_theory/liouville/measure.lean +++ b/src/number_theory/liouville/measure.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Yury G. Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury G. Kudryashov -/ -import measure_theory.measure.lebesgue +import measure_theory.measure.lebesgue.basic import number_theory.liouville.residual import number_theory.liouville.liouville_with import analysis.p_series diff --git a/src/probability/density.lean b/src/probability/density.lean index 2e622d3848004..e4c992f80cc1c 100644 --- a/src/probability/density.lean +++ b/src/probability/density.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kexing Ying -/ import measure_theory.decomposition.radon_nikodym -import measure_theory.measure.haar_of_basis +import measure_theory.measure.haar.of_basis /-! # Probability density function diff --git a/src/probability/kernel/basic.lean b/src/probability/kernel/basic.lean index e2068f99f01cc..9bef08cbd895f 100644 --- a/src/probability/kernel/basic.lean +++ b/src/probability/kernel/basic.lean @@ -3,8 +3,8 @@ Copyright (c) 2022 Rémy Degenne. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ - -import measure_theory.constructions.prod +import measure_theory.integral.bochner +import measure_theory.constructions.prod.basic /-! # Markov Kernels diff --git a/test/monotonicity.lean b/test/monotonicity.lean index c4ab5d00dbac9..411625aed3dc9 100644 --- a/test/monotonicity.lean +++ b/test/monotonicity.lean @@ -6,8 +6,9 @@ Authors: Simon Hudon import tactic.monotonicity import tactic.norm_num import algebra.order.ring.defs -import measure_theory.measure.lebesgue import measure_theory.function.locally_integrable +import measure_theory.integral.bochner +import measure_theory.measure.lebesgue.basic import data.list.defs open list tactic tactic.interactive set From 7ae139f966795f684fc689186f9ccbaedd31bf31 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 19 May 2023 07:37:42 +0000 Subject: [PATCH 1038/1291] feat(linear_algebra/matrix/charpoly/minpoly): `minpoly` is unaffected by `to_lin` and `to_matrix` (#19036) --- .../matrix/charpoly/minpoly.lean | 30 ++++++++++++++++--- 1 file changed, 26 insertions(+), 4 deletions(-) diff --git a/src/linear_algebra/matrix/charpoly/minpoly.lean b/src/linear_algebra/matrix/charpoly/minpoly.lean index 797e0cd6d565d..f8d20bd6c3433 100644 --- a/src/linear_algebra/matrix/charpoly/minpoly.lean +++ b/src/linear_algebra/matrix/charpoly/minpoly.lean @@ -1,7 +1,7 @@ /- 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 +Authors: Aaron Anderson, Jalex Stark, Eric Wieser -/ import linear_algebra.matrix.charpoly.coeff @@ -10,22 +10,32 @@ import ring_theory.power_basis /-! # The minimal polynomial divides the characteristic polynomial of a matrix. + +This also includes some miscellaneous results about `minpoly` on matrices. -/ noncomputable theory -universes u v +universes u v w open polynomial matrix variables {R : Type u} [comm_ring R] variables {n : Type v} [decidable_eq n] [fintype n] +variables {N : Type w} [add_comm_group N] [module R N] open finset -variable {M : matrix n n R} - namespace matrix +open_locale matrix +variables (M : matrix n n R) + +@[simp] theorem minpoly_to_lin' : minpoly R M.to_lin' = minpoly R M := +minpoly.minpoly_alg_equiv (to_lin_alg_equiv' : matrix n n R ≃ₐ[R] _) M + +@[simp] theorem minpoly_to_lin (b : basis n R N) (M : matrix n n R) : + minpoly R (to_lin b b M) = minpoly R M := +minpoly.minpoly_alg_equiv (to_lin_alg_equiv b : matrix n n R ≃ₐ[R] _) M theorem is_integral : is_integral R M := ⟨M.charpoly, ⟨charpoly_monic M, aeval_self_charpoly M⟩⟩ @@ -35,6 +45,18 @@ minpoly.dvd _ _ (aeval_self_charpoly M) end matrix +namespace linear_map + +@[simp] theorem minpoly_to_matrix' (f : (n → R) →ₗ[R] (n → R)) : + minpoly R f.to_matrix' = minpoly R f := +minpoly.minpoly_alg_equiv (to_matrix_alg_equiv' : _ ≃ₐ[R] matrix n n R) f + +@[simp] theorem minpoly_to_matrix (b : basis n R N) (f : N →ₗ[R] N) : + minpoly R (to_matrix b b f) = minpoly R f := +minpoly.minpoly_alg_equiv (to_matrix_alg_equiv b : _ ≃ₐ[R] matrix n n R) f + +end linear_map + section power_basis open algebra From 915591b2bb3ea303648db07284a161a7f2a9e3d4 Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Fri, 19 May 2023 12:01:13 +0000 Subject: [PATCH 1039/1291] chore(analysis/convex/cone/basic): split (#19043) Split out the inner product space material (the dual cone) from `analysis/convex/cone/basic`. What's left imports almost nothing, and can probably be ported immediately. --- src/analysis/convex/cone/basic.lean | 193 +--------------- src/analysis/convex/cone/dual.lean | 210 ++++++++++++++++++ src/analysis/convex/cone/proper.lean | 3 +- .../normed_space/hahn_banach/extension.lean | 1 + .../normed_space/hahn_banach/separation.lean | 1 + .../function/ae_eq_of_integral.lean | 1 + 6 files changed, 219 insertions(+), 190 deletions(-) create mode 100644 src/analysis/convex/cone/dual.lean diff --git a/src/analysis/convex/cone/basic.lean b/src/analysis/convex/cone/basic.lean index ca62b0911010e..ef9525cbd03ff 100644 --- a/src/analysis/convex/cone/basic.lean +++ b/src/analysis/convex/cone/basic.lean @@ -1,9 +1,11 @@ /- -Copyright (c) 2020 Yury Kudryashov All rights reserved. +Copyright (c) 2020 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov, Frédéric Dupuis -/ -import analysis.inner_product_space.projection +import analysis.convex.hull +import data.real.basic +import linear_algebra.linear_pmap /-! # Convex cones @@ -17,9 +19,6 @@ convex cones and ordered modules. We define `convex.to_cone` to be the minimal cone that includes a given convex set. -We define `set.inner_dual_cone` to be the cone consisting of all points `y` such that for -all points `x` in a given set `0 ≤ ⟪ x, y ⟫`. - ## Main statements We prove two extension theorems: @@ -45,7 +44,6 @@ We prove the following theorems: interpretation of the [Farkas lemma](https://en.wikipedia.org/wiki/Farkas%27_lemma#Geometric_interpretation). * `convex_cone.inner_dual_cone_of_inner_dual_cone_eq_self`: - The `inner_dual_cone` of the `inner_dual_cone` of a nonempty, closed, convex cone is itself. ## Implementation notes @@ -58,6 +56,7 @@ While `convex 𝕜` is a predicate on sets, `convex_cone 𝕜 E` is a bundled co * [Emo Welzl and Bernd Gärtner, *Cone Programming*][welzl_garter] -/ +assert_not_exists normed_space open set linear_map open_locale classical pointwise @@ -713,185 +712,3 @@ begin simp only [convex_cone.mem_mk, mem_set_of_eq, subtype.coe_mk, prod.fst_add, prod.snd_add, zero_add, sub_add_cancel] } end - -/-! ### The dual cone -/ - -section dual -variables {H : Type*} [normed_add_comm_group H] [inner_product_space ℝ H] (s t : set H) -open_locale real_inner_product_space - -/-- The dual cone is the cone consisting of all points `y` such that for -all points `x` in a given set `0 ≤ ⟪ x, y ⟫`. -/ -def set.inner_dual_cone (s : set H) : convex_cone ℝ H := -{ carrier := { y | ∀ x ∈ s, 0 ≤ ⟪ x, y ⟫ }, - smul_mem' := λ c hc y hy x hx, - begin - rw real_inner_smul_right, - exact mul_nonneg hc.le (hy x hx) - end, - add_mem' := λ u hu v hv x hx, - begin - rw inner_add_right, - exact add_nonneg (hu x hx) (hv x hx) - end } - -@[simp] lemma mem_inner_dual_cone (y : H) (s : set H) : - y ∈ s.inner_dual_cone ↔ ∀ x ∈ s, 0 ≤ ⟪ x, y ⟫ := iff.rfl - -@[simp] lemma inner_dual_cone_empty : (∅ : set H).inner_dual_cone = ⊤ := -eq_top_iff.mpr $ λ x hy y, false.elim - -/-- Dual cone of the convex cone {0} is the total space. -/ -@[simp] lemma inner_dual_cone_zero : (0 : set H).inner_dual_cone = ⊤ := -eq_top_iff.mpr $ λ x hy y (hy : y = 0), hy.symm ▸ (inner_zero_left _).ge - -/-- Dual cone of the total space is the convex cone {0}. -/ -@[simp] lemma inner_dual_cone_univ : (univ : set H).inner_dual_cone = 0 := -begin - suffices : ∀ x : H, x ∈ (univ : set H).inner_dual_cone → x = 0, - { apply set_like.coe_injective, - exact eq_singleton_iff_unique_mem.mpr ⟨λ x hx, (inner_zero_right _).ge, this⟩ }, - exact λ x hx, by simpa [←real_inner_self_nonpos] using hx (-x) (mem_univ _), -end - -lemma inner_dual_cone_le_inner_dual_cone (h : t ⊆ s) : - s.inner_dual_cone ≤ t.inner_dual_cone := -λ y hy x hx, hy x (h hx) - -lemma pointed_inner_dual_cone : s.inner_dual_cone.pointed := -λ x hx, by rw inner_zero_right - -/-- The inner dual cone of a singleton is given by the preimage of the positive cone under the -linear map `λ y, ⟪x, y⟫`. -/ -lemma inner_dual_cone_singleton (x : H) : - ({x} : set H).inner_dual_cone = (convex_cone.positive ℝ ℝ).comap (innerₛₗ ℝ x) := -convex_cone.ext $ λ i, forall_eq - -lemma inner_dual_cone_union (s t : set H) : - (s ∪ t).inner_dual_cone = s.inner_dual_cone ⊓ t.inner_dual_cone := -le_antisymm - (le_inf (λ x hx y hy, hx _ $ or.inl hy) (λ x hx y hy, hx _ $ or.inr hy)) - (λ x hx y, or.rec (hx.1 _) (hx.2 _)) - -lemma inner_dual_cone_insert (x : H) (s : set H) : - (insert x s).inner_dual_cone = set.inner_dual_cone {x} ⊓ s.inner_dual_cone := -by rw [insert_eq, inner_dual_cone_union] - -lemma inner_dual_cone_Union {ι : Sort*} (f : ι → set H) : - (⋃ i, f i).inner_dual_cone = ⨅ i, (f i).inner_dual_cone := -begin - refine le_antisymm (le_infi $ λ i x hx y hy, hx _ $ mem_Union_of_mem _ hy) _, - intros x hx y hy, - rw [convex_cone.mem_infi] at hx, - obtain ⟨j, hj⟩ := mem_Union.mp hy, - exact hx _ _ hj, -end - -lemma inner_dual_cone_sUnion (S : set (set H)) : - (⋃₀ S).inner_dual_cone = Inf (set.inner_dual_cone '' S) := -by simp_rw [Inf_image, sUnion_eq_bUnion, inner_dual_cone_Union] - -/-- The dual cone of `s` equals the intersection of dual cones of the points in `s`. -/ -lemma inner_dual_cone_eq_Inter_inner_dual_cone_singleton : - (s.inner_dual_cone : set H) = ⋂ i : s, (({i} : set H).inner_dual_cone : set H) := -by rw [←convex_cone.coe_infi, ←inner_dual_cone_Union, Union_of_singleton_coe] - -lemma is_closed_inner_dual_cone : is_closed (s.inner_dual_cone : set H) := -begin - -- reduce the problem to showing that dual cone of a singleton `{x}` is closed - rw inner_dual_cone_eq_Inter_inner_dual_cone_singleton, - apply is_closed_Inter, - intros x, - - -- the dual cone of a singleton `{x}` is the preimage of `[0, ∞)` under `inner x` - have h : ↑({x} : set H).inner_dual_cone = (inner x : H → ℝ) ⁻¹' set.Ici 0, - { rw [inner_dual_cone_singleton, convex_cone.coe_comap, convex_cone.coe_positive, - innerₛₗ_apply_coe] }, - - -- the preimage is closed as `inner x` is continuous and `[0, ∞)` is closed - rw h, - exact is_closed_Ici.preimage (by continuity), -end - -lemma convex_cone.pointed_of_nonempty_of_is_closed (K : convex_cone ℝ H) - (ne : (K : set H).nonempty) (hc : is_closed (K : set H)) : K.pointed := -begin - obtain ⟨x, hx⟩ := ne, - let f : ℝ → H := (• x), - - -- f (0, ∞) is a subset of K - have fI : f '' set.Ioi 0 ⊆ (K : set H), - { rintro _ ⟨_, h, rfl⟩, - exact K.smul_mem (set.mem_Ioi.1 h) hx }, - - -- closure of f (0, ∞) is a subset of K - have clf : closure (f '' set.Ioi 0) ⊆ (K : set H) := hc.closure_subset_iff.2 fI, - - -- f is continuous at 0 from the right - have fc : continuous_within_at f (set.Ioi (0 : ℝ)) 0 := - (continuous_id.smul continuous_const).continuous_within_at, - - -- 0 belongs to the closure of the f (0, ∞) - have mem₀ := fc.mem_closure_image (by rw [closure_Ioi (0 : ℝ), mem_Ici]), - - -- as 0 ∈ closure f (0, ∞) and closure f (0, ∞) ⊆ K, 0 ∈ K. - have f₀ : f 0 = 0 := zero_smul ℝ x, - simpa only [f₀, convex_cone.pointed, ← set_like.mem_coe] using mem_of_subset_of_mem clf mem₀, -end - -section complete_space -variables [complete_space H] - -/-- This is a stronger version of the Hahn-Banach separation theorem for closed convex cones. This -is also the geometric interpretation of Farkas' lemma. -/ -theorem convex_cone.hyperplane_separation_of_nonempty_of_is_closed_of_nmem (K : convex_cone ℝ H) - (ne : (K : set H).nonempty) (hc : is_closed (K : set H)) {b : H} (disj : b ∉ K) : - ∃ (y : H), (∀ x : H, x ∈ K → 0 ≤ ⟪x, y⟫_ℝ) ∧ ⟪y, b⟫_ℝ < 0 := -begin - -- let `z` be the point in `K` closest to `b` - obtain ⟨z, hzK, infi⟩ := exists_norm_eq_infi_of_complete_convex ne hc.is_complete K.convex b, - - -- for any `w` in `K`, we have `⟪b - z, w - z⟫_ℝ ≤ 0` - have hinner := (norm_eq_infi_iff_real_inner_le_zero K.convex hzK).1 infi, - - -- set `y := z - b` - use z - b, - - split, - { -- the rest of the proof is a straightforward calculation - rintros x hxK, - specialize hinner _ (K.add_mem hxK hzK), - rwa [add_sub_cancel, real_inner_comm, ← neg_nonneg, neg_eq_neg_one_mul, - ← real_inner_smul_right, neg_smul, one_smul, neg_sub] at hinner }, - { -- as `K` is closed and non-empty, it is pointed - have hinner₀ := hinner 0 (K.pointed_of_nonempty_of_is_closed ne hc), - - -- the rest of the proof is a straightforward calculation - rw [zero_sub, inner_neg_right, right.neg_nonpos_iff] at hinner₀, - have hbz : b - z ≠ 0 := by { rw sub_ne_zero, contrapose! hzK, rwa ← hzK }, - rw [← neg_zero, lt_neg, ← neg_one_mul, ← real_inner_smul_left, smul_sub, neg_smul, one_smul, - neg_smul, neg_sub_neg, one_smul], - calc 0 < ⟪b - z, b - z⟫_ℝ : lt_of_not_le ((iff.not real_inner_self_nonpos).2 hbz) - ... = ⟪b - z, b - z⟫_ℝ + 0 : (add_zero _).symm - ... ≤ ⟪b - z, b - z⟫_ℝ + ⟪b - z, z⟫_ℝ : add_le_add rfl.ge hinner₀ - ... = ⟪b - z, b - z + z⟫_ℝ : (inner_add_right _ _ _).symm - ... = ⟪b - z, b⟫_ℝ : by rw sub_add_cancel }, -end - -/-- The inner dual of inner dual of a non-empty, closed convex cone is itself. -/ -theorem convex_cone.inner_dual_cone_of_inner_dual_cone_eq_self (K : convex_cone ℝ H) - (ne : (K : set H).nonempty) (hc : is_closed (K : set H)) : - ((K : set H).inner_dual_cone : set H).inner_dual_cone = K := -begin - ext x, - split, - { rw [mem_inner_dual_cone, ← set_like.mem_coe], - contrapose!, - exact K.hyperplane_separation_of_nonempty_of_is_closed_of_nmem ne hc }, - { rintro hxK y h, - specialize h x hxK, - rwa real_inner_comm }, -end - -end complete_space -end dual diff --git a/src/analysis/convex/cone/dual.lean b/src/analysis/convex/cone/dual.lean new file mode 100644 index 0000000000000..a80d36f8cff61 --- /dev/null +++ b/src/analysis/convex/cone/dual.lean @@ -0,0 +1,210 @@ +/- +Copyright (c) 2021 Alexander Bentkamp. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Alexander Bentkamp +-/ +import analysis.convex.cone.basic +import analysis.inner_product_space.projection + +/-! +# Convex cones in inner product spaces + +We define `set.inner_dual_cone` to be the cone consisting of all points `y` such that for +all points `x` in a given set `0 ≤ ⟪ x, y ⟫`. + +## Main statements + +We prove the following theorems: +* `convex_cone.inner_dual_cone_of_inner_dual_cone_eq_self`: + The `inner_dual_cone` of the `inner_dual_cone` of a nonempty, closed, convex cone is itself. + +-/ + + +open set linear_map +open_locale classical pointwise + +variables {𝕜 E F G : Type*} + + +/-! ### The dual cone -/ + +section dual +variables {H : Type*} [normed_add_comm_group H] [inner_product_space ℝ H] (s t : set H) +open_locale real_inner_product_space + +/-- The dual cone is the cone consisting of all points `y` such that for +all points `x` in a given set `0 ≤ ⟪ x, y ⟫`. -/ +def set.inner_dual_cone (s : set H) : convex_cone ℝ H := +{ carrier := { y | ∀ x ∈ s, 0 ≤ ⟪ x, y ⟫ }, + smul_mem' := λ c hc y hy x hx, + begin + rw real_inner_smul_right, + exact mul_nonneg hc.le (hy x hx) + end, + add_mem' := λ u hu v hv x hx, + begin + rw inner_add_right, + exact add_nonneg (hu x hx) (hv x hx) + end } + +@[simp] lemma mem_inner_dual_cone (y : H) (s : set H) : + y ∈ s.inner_dual_cone ↔ ∀ x ∈ s, 0 ≤ ⟪ x, y ⟫ := iff.rfl + +@[simp] lemma inner_dual_cone_empty : (∅ : set H).inner_dual_cone = ⊤ := +eq_top_iff.mpr $ λ x hy y, false.elim + +/-- Dual cone of the convex cone {0} is the total space. -/ +@[simp] lemma inner_dual_cone_zero : (0 : set H).inner_dual_cone = ⊤ := +eq_top_iff.mpr $ λ x hy y (hy : y = 0), hy.symm ▸ (inner_zero_left _).ge + +/-- Dual cone of the total space is the convex cone {0}. -/ +@[simp] lemma inner_dual_cone_univ : (univ : set H).inner_dual_cone = 0 := +begin + suffices : ∀ x : H, x ∈ (univ : set H).inner_dual_cone → x = 0, + { apply set_like.coe_injective, + exact eq_singleton_iff_unique_mem.mpr ⟨λ x hx, (inner_zero_right _).ge, this⟩ }, + exact λ x hx, by simpa [←real_inner_self_nonpos] using hx (-x) (mem_univ _), +end + +lemma inner_dual_cone_le_inner_dual_cone (h : t ⊆ s) : + s.inner_dual_cone ≤ t.inner_dual_cone := +λ y hy x hx, hy x (h hx) + +lemma pointed_inner_dual_cone : s.inner_dual_cone.pointed := +λ x hx, by rw inner_zero_right + +/-- The inner dual cone of a singleton is given by the preimage of the positive cone under the +linear map `λ y, ⟪x, y⟫`. -/ +lemma inner_dual_cone_singleton (x : H) : + ({x} : set H).inner_dual_cone = (convex_cone.positive ℝ ℝ).comap (innerₛₗ ℝ x) := +convex_cone.ext $ λ i, forall_eq + +lemma inner_dual_cone_union (s t : set H) : + (s ∪ t).inner_dual_cone = s.inner_dual_cone ⊓ t.inner_dual_cone := +le_antisymm + (le_inf (λ x hx y hy, hx _ $ or.inl hy) (λ x hx y hy, hx _ $ or.inr hy)) + (λ x hx y, or.rec (hx.1 _) (hx.2 _)) + +lemma inner_dual_cone_insert (x : H) (s : set H) : + (insert x s).inner_dual_cone = set.inner_dual_cone {x} ⊓ s.inner_dual_cone := +by rw [insert_eq, inner_dual_cone_union] + +lemma inner_dual_cone_Union {ι : Sort*} (f : ι → set H) : + (⋃ i, f i).inner_dual_cone = ⨅ i, (f i).inner_dual_cone := +begin + refine le_antisymm (le_infi $ λ i x hx y hy, hx _ $ mem_Union_of_mem _ hy) _, + intros x hx y hy, + rw [convex_cone.mem_infi] at hx, + obtain ⟨j, hj⟩ := mem_Union.mp hy, + exact hx _ _ hj, +end + +lemma inner_dual_cone_sUnion (S : set (set H)) : + (⋃₀ S).inner_dual_cone = Inf (set.inner_dual_cone '' S) := +by simp_rw [Inf_image, sUnion_eq_bUnion, inner_dual_cone_Union] + +/-- The dual cone of `s` equals the intersection of dual cones of the points in `s`. -/ +lemma inner_dual_cone_eq_Inter_inner_dual_cone_singleton : + (s.inner_dual_cone : set H) = ⋂ i : s, (({i} : set H).inner_dual_cone : set H) := +by rw [←convex_cone.coe_infi, ←inner_dual_cone_Union, Union_of_singleton_coe] + +lemma is_closed_inner_dual_cone : is_closed (s.inner_dual_cone : set H) := +begin + -- reduce the problem to showing that dual cone of a singleton `{x}` is closed + rw inner_dual_cone_eq_Inter_inner_dual_cone_singleton, + apply is_closed_Inter, + intros x, + + -- the dual cone of a singleton `{x}` is the preimage of `[0, ∞)` under `inner x` + have h : ↑({x} : set H).inner_dual_cone = (inner x : H → ℝ) ⁻¹' set.Ici 0, + { rw [inner_dual_cone_singleton, convex_cone.coe_comap, convex_cone.coe_positive, + innerₛₗ_apply_coe] }, + + -- the preimage is closed as `inner x` is continuous and `[0, ∞)` is closed + rw h, + exact is_closed_Ici.preimage (by continuity), +end + +lemma convex_cone.pointed_of_nonempty_of_is_closed (K : convex_cone ℝ H) + (ne : (K : set H).nonempty) (hc : is_closed (K : set H)) : K.pointed := +begin + obtain ⟨x, hx⟩ := ne, + let f : ℝ → H := (• x), + + -- f (0, ∞) is a subset of K + have fI : f '' set.Ioi 0 ⊆ (K : set H), + { rintro _ ⟨_, h, rfl⟩, + exact K.smul_mem (set.mem_Ioi.1 h) hx }, + + -- closure of f (0, ∞) is a subset of K + have clf : closure (f '' set.Ioi 0) ⊆ (K : set H) := hc.closure_subset_iff.2 fI, + + -- f is continuous at 0 from the right + have fc : continuous_within_at f (set.Ioi (0 : ℝ)) 0 := + (continuous_id.smul continuous_const).continuous_within_at, + + -- 0 belongs to the closure of the f (0, ∞) + have mem₀ := fc.mem_closure_image (by rw [closure_Ioi (0 : ℝ), mem_Ici]), + + -- as 0 ∈ closure f (0, ∞) and closure f (0, ∞) ⊆ K, 0 ∈ K. + have f₀ : f 0 = 0 := zero_smul ℝ x, + simpa only [f₀, convex_cone.pointed, ← set_like.mem_coe] using mem_of_subset_of_mem clf mem₀, +end + +section complete_space +variables [complete_space H] + +/-- This is a stronger version of the Hahn-Banach separation theorem for closed convex cones. This +is also the geometric interpretation of Farkas' lemma. -/ +theorem convex_cone.hyperplane_separation_of_nonempty_of_is_closed_of_nmem (K : convex_cone ℝ H) + (ne : (K : set H).nonempty) (hc : is_closed (K : set H)) {b : H} (disj : b ∉ K) : + ∃ (y : H), (∀ x : H, x ∈ K → 0 ≤ ⟪x, y⟫_ℝ) ∧ ⟪y, b⟫_ℝ < 0 := +begin + -- let `z` be the point in `K` closest to `b` + obtain ⟨z, hzK, infi⟩ := exists_norm_eq_infi_of_complete_convex ne hc.is_complete K.convex b, + + -- for any `w` in `K`, we have `⟪b - z, w - z⟫_ℝ ≤ 0` + have hinner := (norm_eq_infi_iff_real_inner_le_zero K.convex hzK).1 infi, + + -- set `y := z - b` + use z - b, + + split, + { -- the rest of the proof is a straightforward calculation + rintros x hxK, + specialize hinner _ (K.add_mem hxK hzK), + rwa [add_sub_cancel, real_inner_comm, ← neg_nonneg, neg_eq_neg_one_mul, + ← real_inner_smul_right, neg_smul, one_smul, neg_sub] at hinner }, + { -- as `K` is closed and non-empty, it is pointed + have hinner₀ := hinner 0 (K.pointed_of_nonempty_of_is_closed ne hc), + + -- the rest of the proof is a straightforward calculation + rw [zero_sub, inner_neg_right, right.neg_nonpos_iff] at hinner₀, + have hbz : b - z ≠ 0 := by { rw sub_ne_zero, contrapose! hzK, rwa ← hzK }, + rw [← neg_zero, lt_neg, ← neg_one_mul, ← real_inner_smul_left, smul_sub, neg_smul, one_smul, + neg_smul, neg_sub_neg, one_smul], + calc 0 < ⟪b - z, b - z⟫_ℝ : lt_of_not_le ((iff.not real_inner_self_nonpos).2 hbz) + ... = ⟪b - z, b - z⟫_ℝ + 0 : (add_zero _).symm + ... ≤ ⟪b - z, b - z⟫_ℝ + ⟪b - z, z⟫_ℝ : add_le_add rfl.ge hinner₀ + ... = ⟪b - z, b - z + z⟫_ℝ : (inner_add_right _ _ _).symm + ... = ⟪b - z, b⟫_ℝ : by rw sub_add_cancel }, +end + +/-- The inner dual of inner dual of a non-empty, closed convex cone is itself. -/ +theorem convex_cone.inner_dual_cone_of_inner_dual_cone_eq_self (K : convex_cone ℝ H) + (ne : (K : set H).nonempty) (hc : is_closed (K : set H)) : + ((K : set H).inner_dual_cone : set H).inner_dual_cone = K := +begin + ext x, + split, + { rw [mem_inner_dual_cone, ← set_like.mem_coe], + contrapose!, + exact K.hyperplane_separation_of_nonempty_of_is_closed_of_nmem ne hc }, + { rintro hxK y h, + specialize h x hxK, + rwa real_inner_comm }, +end + +end complete_space +end dual diff --git a/src/analysis/convex/cone/proper.lean b/src/analysis/convex/cone/proper.lean index 96ee065db043c..e210123b9d5e6 100644 --- a/src/analysis/convex/cone/proper.lean +++ b/src/analysis/convex/cone/proper.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Apurva Nakade -/ import analysis.convex.cone.basic +import topology.algebra.monoid /-! @@ -29,8 +30,6 @@ The next steps are: -/ -open continuous_linear_map filter - namespace convex_cone variables {E : Type*} [add_comm_monoid E] [has_smul ℝ E] [topological_space E] diff --git a/src/analysis/normed_space/hahn_banach/extension.lean b/src/analysis/normed_space/hahn_banach/extension.lean index 80ab1610e970d..f4c581a7f318e 100644 --- a/src/analysis/normed_space/hahn_banach/extension.lean +++ b/src/analysis/normed_space/hahn_banach/extension.lean @@ -6,6 +6,7 @@ Authors: Yury Kudryashov, Heather Macbeth import analysis.convex.cone.basic import analysis.normed_space.is_R_or_C import analysis.normed_space.extend +import data.is_R_or_C.lemmas /-! # Extension Hahn-Banach theorem diff --git a/src/analysis/normed_space/hahn_banach/separation.lean b/src/analysis/normed_space/hahn_banach/separation.lean index 615487672eb2b..ec6836e48ff2b 100644 --- a/src/analysis/normed_space/hahn_banach/separation.lean +++ b/src/analysis/normed_space/hahn_banach/separation.lean @@ -5,6 +5,7 @@ Authors: Bhavik Mehta, Yaël Dillies -/ import analysis.convex.cone.basic import analysis.convex.gauge +import topology.algebra.module.finite_dimension import topology.algebra.module.locally_convex /-! diff --git a/src/measure_theory/function/ae_eq_of_integral.lean b/src/measure_theory/function/ae_eq_of_integral.lean index 9ab0cfcf0be01..6358a66f7fa7d 100644 --- a/src/measure_theory/function/ae_eq_of_integral.lean +++ b/src/measure_theory/function/ae_eq_of_integral.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ +import analysis.inner_product_space.basic import analysis.normed_space.dual import measure_theory.function.strongly_measurable.lp import measure_theory.integral.set_integral From a7c017d750512a352b623b1824d75da5998457d0 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Fri, 19 May 2023 13:16:23 +0000 Subject: [PATCH 1040/1291] chore(ring_theory/localization/away): split (#19041) This breaks off a large initial segment of the [longest chain](https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/port.20progress/near/359494005) remaining to port. Co-authored-by: Scott Morrison --- src/algebra/category/Ring/instances.lean | 2 +- .../prime_spectrum/basic.lean | 2 +- src/ring_theory/jacobson.lean | 2 +- src/ring_theory/local_properties.lean | 2 +- .../localization/away/adjoin_root.lean | 37 +++++++++++++++++++ .../{away.lean => away/basic.lean} | 27 ++------------ src/ring_theory/ring_hom/integral.lean | 1 + src/ring_theory/ring_hom_properties.lean | 2 +- 8 files changed, 46 insertions(+), 29 deletions(-) create mode 100644 src/ring_theory/localization/away/adjoin_root.lean rename src/ring_theory/localization/{away.lean => away/basic.lean} (90%) diff --git a/src/algebra/category/Ring/instances.lean b/src/algebra/category/Ring/instances.lean index b224997a96e49..cef4e306a2150 100644 --- a/src/algebra/category/Ring/instances.lean +++ b/src/algebra/category/Ring/instances.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Andrew Yang -/ import algebra.category.Ring.basic -import ring_theory.localization.away +import ring_theory.localization.away.basic import ring_theory.ideal.local_ring /-! diff --git a/src/algebraic_geometry/prime_spectrum/basic.lean b/src/algebraic_geometry/prime_spectrum/basic.lean index ad98e61dcd8fd..559f61d56941d 100644 --- a/src/algebraic_geometry/prime_spectrum/basic.lean +++ b/src/algebraic_geometry/prime_spectrum/basic.lean @@ -7,7 +7,7 @@ import algebra.punit_instances import linear_algebra.finsupp import ring_theory.ideal.over import ring_theory.ideal.prod -import ring_theory.localization.away +import ring_theory.localization.away.basic import ring_theory.nilpotent import topology.sets.closeds import topology.sober diff --git a/src/ring_theory/jacobson.lean b/src/ring_theory/jacobson.lean index c4f252cd949d8..a44c12838c5c5 100644 --- a/src/ring_theory/jacobson.lean +++ b/src/ring_theory/jacobson.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 Devon Tuma. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Devon Tuma -/ -import ring_theory.localization.away +import ring_theory.localization.away.basic import ring_theory.ideal.over import ring_theory.jacobson_ideal diff --git a/src/ring_theory/local_properties.lean b/src/ring_theory/local_properties.lean index 70b0e5931f852..c958abad5cea5 100644 --- a/src/ring_theory/local_properties.lean +++ b/src/ring_theory/local_properties.lean @@ -5,7 +5,7 @@ Authors: Andrew Yang -/ import ring_theory.finite_type import ring_theory.localization.at_prime -import ring_theory.localization.away +import ring_theory.localization.away.basic import ring_theory.localization.integer import ring_theory.localization.submodule import ring_theory.nilpotent diff --git a/src/ring_theory/localization/away/adjoin_root.lean b/src/ring_theory/localization/away/adjoin_root.lean new file mode 100644 index 0000000000000..c169a02f9172c --- /dev/null +++ b/src/ring_theory/localization/away/adjoin_root.lean @@ -0,0 +1,37 @@ +/- +Copyright (c) 2018 Kenny Lau. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Kenny Lau, Mario Carneiro, Johan Commelin, Amelia Livingston, Anne Baanen +-/ +import ring_theory.adjoin_root +import ring_theory.localization.away.basic + +/-! +The `R`-`alg_equiv` between the localization of `R` away from `r` and +`R` with an inverse of `r` adjoined. +-/ + +open polynomial adjoin_root localization + +variables {R : Type*} [comm_ring R] + +local attribute [instance] is_localization.alg_hom_subsingleton adjoin_root.alg_hom_subsingleton + +/-- The `R`-`alg_equiv` between the localization of `R` away from `r` and + `R` with an inverse of `r` adjoined. -/ +noncomputable def localization.away_equiv_adjoin (r : R) : away r ≃ₐ[R] adjoin_root (C r * X - 1) := +alg_equiv.of_alg_hom + { commutes' := is_localization.away.away_map.lift_eq r + (is_unit_of_mul_eq_one _ _ $ root_is_inv r), .. away_lift _ r _ } + (lift_hom _ (is_localization.away.inv_self r) $ by simp only + [map_sub, map_mul, aeval_C, aeval_X, is_localization.away.mul_inv_self, aeval_one, sub_self]) + (subsingleton.elim _ _) + (subsingleton.elim _ _) + +lemma is_localization.adjoin_inv (r : R) : is_localization.away r (adjoin_root $ C r * X - 1) := +is_localization.is_localization_of_alg_equiv _ (localization.away_equiv_adjoin r) + +lemma is_localization.away.finite_presentation (r : R) {S} [comm_ring S] [algebra R S] + [is_localization.away r S] : algebra.finite_presentation R S := +(adjoin_root.finite_presentation _).equiv $ (localization.away_equiv_adjoin r).symm.trans $ + is_localization.alg_equiv (submonoid.powers r) _ _ diff --git a/src/ring_theory/localization/away.lean b/src/ring_theory/localization/away/basic.lean similarity index 90% rename from src/ring_theory/localization/away.lean rename to src/ring_theory/localization/away/basic.lean index 0ea25713791eb..bcdfee9da33f6 100644 --- a/src/ring_theory/localization/away.lean +++ b/src/ring_theory/localization/away/basic.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, Mario Carneiro, Johan Commelin, Amelia Livingston, Anne Baanen -/ -import ring_theory.adjoin_root +import ring_theory.unique_factorization_domain import ring_theory.localization.basic /-! @@ -172,34 +172,13 @@ end localization end comm_semiring -open polynomial adjoin_root localization +open localization variables {R : Type*} [comm_ring R] -local attribute [instance] is_localization.alg_hom_subsingleton adjoin_root.alg_hom_subsingleton - -/-- The `R`-`alg_equiv` between the localization of `R` away from `r` and - `R` with an inverse of `r` adjoined. -/ -noncomputable def localization.away_equiv_adjoin (r : R) : away r ≃ₐ[R] adjoin_root (C r * X - 1) := -alg_equiv.of_alg_hom - { commutes' := is_localization.away.away_map.lift_eq r - (is_unit_of_mul_eq_one _ _ $ root_is_inv r), .. away_lift _ r _ } - (lift_hom _ (is_localization.away.inv_self r) $ by simp only - [map_sub, map_mul, aeval_C, aeval_X, is_localization.away.mul_inv_self, aeval_one, sub_self]) - (subsingleton.elim _ _) - (subsingleton.elim _ _) - -lemma is_localization.adjoin_inv (r : R) : is_localization.away r (adjoin_root $ C r * X - 1) := -is_localization.is_localization_of_alg_equiv _ (localization.away_equiv_adjoin r) - -lemma is_localization.away.finite_presentation (r : R) {S} [comm_ring S] [algebra R S] - [is_localization.away r S] : algebra.finite_presentation R S := -(adjoin_root.finite_presentation _).equiv $ (localization.away_equiv_adjoin r).symm.trans $ - is_localization.alg_equiv (submonoid.powers r) _ _ - section num_denom -open multiplicity unique_factorization_monoid is_localization +open unique_factorization_monoid is_localization variable (x : R) diff --git a/src/ring_theory/ring_hom/integral.lean b/src/ring_theory/ring_hom/integral.lean index 09c56649a2549..22318a63766fa 100644 --- a/src/ring_theory/ring_hom/integral.lean +++ b/src/ring_theory/ring_hom/integral.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Andrew Yang -/ import ring_theory.ring_hom_properties +import ring_theory.integral_closure /-! diff --git a/src/ring_theory/ring_hom_properties.lean b/src/ring_theory/ring_hom_properties.lean index 5799d9d7ea037..321ccee362f0b 100644 --- a/src/ring_theory/ring_hom_properties.lean +++ b/src/ring_theory/ring_hom_properties.lean @@ -6,7 +6,7 @@ Authors: Andrew Yang import algebra.category.Ring.constructions import algebra.category.Ring.colimits import category_theory.isomorphism -import ring_theory.localization.away +import ring_theory.localization.away.basic import ring_theory.is_tensor_product /-! From ef7403fe253f8992d57109b63b9b2a0cc44c82cc Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Fri, 19 May 2023 13:16:24 +0000 Subject: [PATCH 1041/1291] chore(.github/workflows): don't rely on linkifier in maintainer merge messages (#19046) Not tested, and I'm just guessing that the relevant properties exist as I can't find the relevant API docs... Co-authored-by: Scott Morrison Co-authored-by: Eric Wieser --- .github/workflows/maintainer_merge_comment.yml | 2 +- .github/workflows/maintainer_merge_review.yml | 2 +- .github/workflows/maintainer_merge_review_comment.yml | 2 +- 3 files changed, 3 insertions(+), 3 deletions(-) diff --git a/.github/workflows/maintainer_merge_comment.yml b/.github/workflows/maintainer_merge_comment.yml index b49bb861f2556..d691f2841e484 100644 --- a/.github/workflows/maintainer_merge_comment.yml +++ b/.github/workflows/maintainer_merge_comment.yml @@ -30,7 +30,7 @@ jobs: type: 'stream' topic: 'maintainer merge' content: | - ${{ format('{0} requested a maintainer merge on PR #{1}:', github.event.comment.user.login, github.event.issue.number) }} + ${{ format('{0} requested a maintainer merge on PR [#{1}]({2}):', github.event.comment.user.login, github.event.issue.number, github.event.issue.html_url) }} > ${{ github.event.issue.title }} diff --git a/.github/workflows/maintainer_merge_review.yml b/.github/workflows/maintainer_merge_review.yml index 5003f2f902825..21cc34de067ea 100644 --- a/.github/workflows/maintainer_merge_review.yml +++ b/.github/workflows/maintainer_merge_review.yml @@ -28,7 +28,7 @@ jobs: type: 'stream' topic: 'maintainer merge' content: | - ${{ format('{0} requested a maintainer merge on PR #{1}:', github.event.review.user.login, github.event.pull_request.number) }} + ${{ format('{0} requested a maintainer merge on PR [#{1}]({2}):', github.event.review.user.login, github.event.pull_request.number, github.event.pull_request.html_url) }} > ${{ github.event.pull_request.title }} diff --git a/.github/workflows/maintainer_merge_review_comment.yml b/.github/workflows/maintainer_merge_review_comment.yml index 2f7f40443e5b4..a3f3b0d627889 100644 --- a/.github/workflows/maintainer_merge_review_comment.yml +++ b/.github/workflows/maintainer_merge_review_comment.yml @@ -28,7 +28,7 @@ jobs: type: 'stream' topic: 'maintainer merge' content: | - ${{ format('{0} requested a maintainer merge on PR #{1}:', github.event.comment.user.login, github.event.pull_request.number) }} + ${{ format('{0} requested a maintainer merge on PR [#{1}]({2}):', github.event.comment.user.login, github.event.pull_request.number, github.event.pull_request.html_url) }} > ${{ github.event.pull_request.title }} From f51de8769c34652d82d1c8e5f8f18f8374782bed Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 19 May 2023 16:48:34 +0000 Subject: [PATCH 1042/1291] feat(combinatorics/simple_graph/regularity/chunk): Partition of a part (#18371) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The heart of the calculation of Szemerédi Regularity Lemma. Define the partition of a part of the original partition and show it locally increases the energy. This is all internal to the proof of SRL, so I made most lemmas `private` --- .../simple_graph/regularity/chunk.lean | 550 ++++++++++++++++++ .../simple_graph/regularity/uniform.lean | 2 +- 2 files changed, 551 insertions(+), 1 deletion(-) create mode 100644 src/combinatorics/simple_graph/regularity/chunk.lean diff --git a/src/combinatorics/simple_graph/regularity/chunk.lean b/src/combinatorics/simple_graph/regularity/chunk.lean new file mode 100644 index 0000000000000..e72a5221269bf --- /dev/null +++ b/src/combinatorics/simple_graph/regularity/chunk.lean @@ -0,0 +1,550 @@ +/- +Copyright (c) 2022 Yaël Dillies, Bhavik Mehta. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies, Bhavik Mehta +-/ +import combinatorics.simple_graph.regularity.bound +import combinatorics.simple_graph.regularity.equitabilise +import combinatorics.simple_graph.regularity.uniform + +/-! +# Chunk of the increment partition for Szemerédi Regularity Lemma + +In the proof of Szemerédi Regularity Lemma, we need to partition each part of a starting partition +to increase the energy. This file defines those partitions of parts and shows that they locally +increase the energy. + +This entire file is internal to the proof of Szemerédi Regularity Lemma. + +## Main declarations + +* `szemeredi_regularity.chunk`: The partition of a part of the starting partition. +* `szemeredi_regularity.edge_density_chunk_uniform`: `chunk` does not locally decrease the edge + density between uniform parts too much. +* `szemeredi_regularity.edge_density_chunk_not_uniform`: `chunk` locally increases the edge density + between non-uniform parts. + +## TODO + +Once ported to mathlib4, this file will be a great golfing ground for Heather's new tactic +`rel_congr`. +-/ + +open finpartition finset fintype rel nat +open_locale big_operators classical + +local attribute [positivity] tactic.positivity_szemeredi_regularity + +namespace szemeredi_regularity +variables {α : Type*} [fintype α] {P : finpartition (univ : finset α)} (hP : P.is_equipartition) + (G : simple_graph α) (ε : ℝ) {U : finset α} (hU : U ∈ P.parts) (V : finset α) + +local notation `m` := (card α/step_bound P.parts.card : ℕ) + +/-! +### Definitions + +We define `chunk`, the partition of a part, and `star`, the sets of parts of `chunk` that are +contained in the corresponding witness of non-uniformity. +-/ + +/-- The portion of `szemeredi_regularity.increment` which partitions `U`. -/ +noncomputable def chunk : finpartition U := +dite (U.card = m * 4 ^ P.parts.card + (card α/P.parts.card - m * 4 ^ P.parts.card)) + (λ hUcard, (atomise U $ P.nonuniform_witnesses G ε U).equitabilise $ card_aux₁ hUcard) + (λ hUcard, (atomise U $ P.nonuniform_witnesses G ε U).equitabilise $ card_aux₂ hP hU hUcard) +-- `hP` and `hU` are used to get that `U` has size +-- `m * 4 ^ P.parts.card + a or m * 4 ^ P.parts.card + a + 1` + +/-- The portion of `szemeredi_regularity.chunk` which is contained in the witness of non uniformity +of `U` and `V`. -/ +noncomputable def star (V : finset α) : finset (finset α) := +(chunk hP G ε hU).parts.filter (⊆ G.nonuniform_witness ε U V) + +/-! +### Density estimates + +We estimate the density between parts of `chunk`. +-/ + +lemma bUnion_star_subset_nonuniform_witness : + (star hP G ε hU V).bUnion id ⊆ G.nonuniform_witness ε U V := +bUnion_subset_iff_forall_subset.2 $ λ A hA, (mem_filter.1 hA).2 + +variables {hP G ε hU V} {𝒜 : finset (finset α)} {s : finset α} + +lemma star_subset_chunk : star hP G ε hU V ⊆ (chunk hP G ε hU).parts := filter_subset _ _ + +private lemma card_nonuniform_witness_sdiff_bUnion_star (hV : V ∈ P.parts) (hUV : U ≠ V) + (h₂ : ¬G.is_uniform ε U V) : + (G.nonuniform_witness ε U V \ (star hP G ε hU V).bUnion id).card ≤ 2 ^ (P.parts.card - 1) * m := +begin + have hX : G.nonuniform_witness ε U V ∈ P.nonuniform_witnesses G ε U := + nonuniform_witness_mem_nonuniform_witnesses h₂ hV hUV, + have q : G.nonuniform_witness ε U V \ (star hP G ε hU V).bUnion id ⊆ + ((atomise U $ P.nonuniform_witnesses G ε U).parts.filter $ + λ B, B ⊆ G.nonuniform_witness ε U V ∧ B.nonempty).bUnion + (λ B, B \ ((chunk hP G ε hU).parts.filter (⊆ B)).bUnion id), + { intros x hx, + rw [←bUnion_filter_atomise hX (G.nonuniform_witness_subset h₂), star, mem_sdiff, mem_bUnion] at + hx, + simp only [not_exists, mem_bUnion, and_imp, filter_congr_decidable, exists_prop, mem_filter, + not_and, mem_sdiff, id.def, mem_sdiff] at hx ⊢, + obtain ⟨⟨B, hB₁, hB₂⟩, hx⟩ := hx, + exact ⟨B, hB₁, hB₂, λ A hA AB, hx A hA $ AB.trans hB₁.2.1⟩ }, + apply (card_le_of_subset q).trans (card_bUnion_le.trans _), + transitivity ∑ i in (atomise U $ P.nonuniform_witnesses G ε U).parts.filter + (λ B, B ⊆ G.nonuniform_witness ε U V ∧ B.nonempty), m, + { suffices : ∀ B ∈ (atomise U $ P.nonuniform_witnesses G ε U).parts, + (B \ ((chunk hP G ε hU).parts.filter (⊆ B)).bUnion id).card ≤ m, + { exact sum_le_sum (λ B hB, this B $ filter_subset _ _ hB) }, + intros B hB, + unfold chunk, + split_ifs with h₁, + { convert card_parts_equitabilise_subset_le _ (card_aux₁ h₁) hB }, + { convert card_parts_equitabilise_subset_le _ (card_aux₂ hP hU h₁) hB } }, + rw sum_const, + refine mul_le_mul_right' _ _, + have t := card_filter_atomise_le_two_pow hX, + rw filter_congr_decidable at t, + refine t.trans (pow_le_pow (by norm_num) $ tsub_le_tsub_right _ _), + exact card_image_le.trans (card_le_of_subset $ filter_subset _ _), +end + +private lemma one_sub_eps_mul_card_nonuniform_witness_le_card_star (hV : V ∈ P.parts) (hUV : U ≠ V) + (hunif : ¬G.is_uniform ε U V) (hPε : 100 ≤ 4 ^ P.parts.card * ε ^ 5) (hε₁ : ε ≤ 1) : + (1 - ε/10) * (G.nonuniform_witness ε U V).card ≤ ((star hP G ε hU V).bUnion id).card := +begin + have hP₁ : 0 < P.parts.card := finset.card_pos.2 ⟨_, hU⟩, + have : (2^P.parts.card : ℝ) * m/(U.card * ε) ≤ ε/10, + { rw [←div_div, div_le_iff'], + swap, + positivity, + refine le_of_mul_le_mul_left _ (pow_pos zero_lt_two P.parts.card), + calc + 2^P.parts.card * ((2^P.parts.card * m : ℝ)/U.card) + = (2 * 2)^P.parts.card * m/U.card : by rw [mul_pow, ←mul_div_assoc, mul_assoc] + ... = 4 ^ P.parts.card * m/U.card : by norm_num + ... ≤ 1 : div_le_one_of_le (pow_mul_m_le_card_part hP hU) (cast_nonneg _) + ... ≤ 2^P.parts.card * ε ^ 2 / 10 : begin + refine (one_le_sq_iff $ by positivity).1 _, + rw [div_pow, mul_pow, pow_right_comm, ←pow_mul ε, + one_le_div (sq_pos_of_ne_zero (10 : ℝ) $ by norm_num)], + calc + (10 ^ 2 : ℝ) + = 100 : by norm_num + ... ≤ 4 ^ P.parts.card * ε ^ 5 : hPε + ... ≤ 4 ^ P.parts.card * ε^4 + : mul_le_mul_of_nonneg_left (pow_le_pow_of_le_one (by positivity) hε₁ $ + le_succ _) (by positivity) + ... = (2 ^ 2)^P.parts.card * ε ^ (2 * 2) : by norm_num, + end + ... = 2^P.parts.card * (ε * (ε / 10)) : by rw [mul_div_assoc, sq, mul_div_assoc] }, + calc + (1 - ε/10) * (G.nonuniform_witness ε U V).card + ≤ (1 - 2^P.parts.card * m/(U.card * ε)) * (G.nonuniform_witness ε U V).card + : mul_le_mul_of_nonneg_right (sub_le_sub_left this _) (cast_nonneg _) + ... = (G.nonuniform_witness ε U V).card - 2^P.parts.card * m/(U.card * ε) + * (G.nonuniform_witness ε U V).card + : by rw [sub_mul, one_mul] + ... ≤ (G.nonuniform_witness ε U V).card - 2^(P.parts.card - 1) * m : begin + refine sub_le_sub_left _ _, + have : (2 : ℝ)^P.parts.card = 2^(P.parts.card - 1) * 2, + { rw [←pow_succ', tsub_add_cancel_of_le (succ_le_iff.2 hP₁)] }, + rw [←mul_div_right_comm, this, mul_right_comm _ (2 : ℝ), mul_assoc, le_div_iff], + refine mul_le_mul_of_nonneg_left _ (by positivity), + exact (G.le_card_nonuniform_witness hunif).trans + (le_mul_of_one_le_left (cast_nonneg _) one_le_two), + have := P.nonempty_of_mem_parts hU, + positivity, + end + ... ≤ ((star hP G ε hU V).bUnion id).card : begin + norm_cast, + rw [sub_le_comm, ←cast_sub (card_le_of_subset $ + bUnion_star_subset_nonuniform_witness hP G ε hU V), ←card_sdiff + (bUnion_star_subset_nonuniform_witness hP G ε hU V), cast_le], + exact card_nonuniform_witness_sdiff_bUnion_star hV hUV hunif, + end +end + +variables {hP G ε U hU V} + +/-! ### `chunk` -/ + +lemma card_chunk (hm : m ≠ 0) : (chunk hP G ε hU).parts.card = 4 ^ P.parts.card := +begin + unfold chunk, + split_ifs, + { rw [card_parts_equitabilise _ _ hm, tsub_add_cancel_of_le], + exact le_of_lt a_add_one_le_four_pow_parts_card }, + { rw [card_parts_equitabilise _ _ hm, tsub_add_cancel_of_le a_add_one_le_four_pow_parts_card] } +end + +lemma card_eq_of_mem_parts_chunk (hs : s ∈ (chunk hP G ε hU).parts) : s.card = m ∨ s.card = m + 1 := +by { unfold chunk at hs, split_ifs at hs; exact card_eq_of_mem_parts_equitabilise hs } + +lemma m_le_card_of_mem_chunk_parts (hs : s ∈ (chunk hP G ε hU).parts) : m ≤ s.card := +(card_eq_of_mem_parts_chunk hs).elim ge_of_eq $ λ i, by simp [i] + +lemma card_le_m_add_one_of_mem_chunk_parts (hs : s ∈ (chunk hP G ε hU).parts) : s.card ≤ m + 1 := +(card_eq_of_mem_parts_chunk hs).elim (λ i, by simp [i]) (λ i, i.le) + +lemma card_bUnion_star_le_m_add_one_card_star_mul : + (((star hP G ε hU V).bUnion id).card : ℝ) ≤ (star hP G ε hU V).card * (m + 1) := +by exact_mod_cast (card_bUnion_le_card_mul _ _ _ $ λ s hs, + card_le_m_add_one_of_mem_chunk_parts $ star_subset_chunk hs) + +private lemma le_sum_card_subset_chunk_parts (h𝒜 : 𝒜 ⊆ (chunk hP G ε hU).parts) (hs : s ∈ 𝒜) : + (𝒜.card : ℝ) * s.card * (m / (m + 1)) ≤ (𝒜.sup id).card := +begin + rw [mul_div_assoc', div_le_iff coe_m_add_one_pos, mul_right_comm], + refine mul_le_mul _ _ (cast_nonneg _) (cast_nonneg _), + { rw [←(of_subset _ h𝒜 rfl).sum_card_parts, of_subset_parts, ←cast_mul, cast_le], + exact card_nsmul_le_sum _ _ _ (λ x hx, m_le_card_of_mem_chunk_parts $ h𝒜 hx) }, + { exact_mod_cast card_le_m_add_one_of_mem_chunk_parts (h𝒜 hs) } +end + +private lemma sum_card_subset_chunk_parts_le (m_pos : (0 : ℝ) < m) + (h𝒜 : 𝒜 ⊆ (chunk hP G ε hU).parts) (hs : s ∈ 𝒜) : + ((𝒜.sup id).card : ℝ) ≤ (𝒜.card * s.card) * ((m+1)/m) := +begin + rw [sup_eq_bUnion, mul_div_assoc', le_div_iff m_pos, mul_right_comm], + refine mul_le_mul _ _ (cast_nonneg _) (by positivity), + { norm_cast, + refine card_bUnion_le_card_mul _ _ _ (λ x hx, _), + apply card_le_m_add_one_of_mem_chunk_parts (h𝒜 hx) }, + { exact_mod_cast m_le_card_of_mem_chunk_parts (h𝒜 hs) } +end + +private lemma one_sub_le_m_div_m_add_one_sq [nonempty α] + (hPα : P.parts.card * 16 ^ P.parts.card ≤ card α) (hPε : 100 ≤ 4 ^ P.parts.card * ε ^ 5) : + 1 - ε ^ 5/50 ≤ (m/(m + 1)) ^ 2 := +begin + have : ((m:ℝ) / (m+1)) = 1 - 1/(m+1), + { rw [one_sub_div coe_m_add_one_pos.ne', add_sub_cancel] }, + rw [this, sub_sq, one_pow, mul_one], + refine le_trans _ (le_add_of_nonneg_right $ sq_nonneg _), + rw [sub_le_sub_iff_left, ←le_div_iff' (show (0:ℝ) < 2, by norm_num), div_div, + one_div_le coe_m_add_one_pos, one_div_div], + refine le_trans _ (le_add_of_nonneg_right zero_le_one), + norm_num, + apply hundred_div_ε_pow_five_le_m hPα hPε, + positivity, +end + +private lemma m_add_one_div_m_le_one_add [nonempty α] + (hPα : P.parts.card * 16 ^ P.parts.card ≤ card α) (hPε : 100 ≤ 4 ^ P.parts.card * ε ^ 5) + (hε₁ : ε ≤ 1) : + ((m + 1 : ℝ)/m) ^ 2 ≤ 1 + ε ^ 5/49 := +begin + rw same_add_div, + swap, + { positivity }, + have : 1 + 1/(m:ℝ) ≤ 1 + ε ^ 5/100, + { rw [add_le_add_iff_left, ←one_div_div (100:ℝ)], + exact one_div_le_one_div_of_le (by positivity) (hundred_div_ε_pow_five_le_m hPα hPε) }, + refine (pow_le_pow_of_le_left _ this 2).trans _, + { positivity }, + rw [add_sq, one_pow, add_assoc, add_le_add_iff_left, mul_one, ←le_sub_iff_add_le', + div_eq_mul_one_div _ (49:ℝ), mul_div_left_comm (2:ℝ), ←mul_sub_left_distrib, div_pow, + div_le_iff (show (0:ℝ) < 100 ^ 2, by norm_num), mul_assoc, sq], + refine mul_le_mul_of_nonneg_left _ (by positivity), + exact (pow_le_one 5 (by positivity) hε₁).trans (by norm_num), +end + +private lemma density_sub_eps_le_sum_density_div_card [nonempty α] + (hPα : P.parts.card * 16 ^ P.parts.card ≤ card α) (hPε : 100 ≤ 4 ^ P.parts.card * ε ^ 5) + {hU : U ∈ P.parts} {hV : V ∈ P.parts} {A B : finset (finset α)} + (hA : A ⊆ (chunk hP G ε hU).parts) (hB : B ⊆ (chunk hP G ε hV).parts) : + ↑(G.edge_density (A.bUnion id) (B.bUnion id)) - ε ^ 5/50 ≤ + (∑ ab in A.product B, G.edge_density ab.1 ab.2)/(A.card * B.card) := +begin + have : ↑(G.edge_density (A.bUnion id) (B.bUnion id)) - ε ^ 5/50 + ≤ (1 - ε ^ 5/50) * G.edge_density (A.bUnion id) (B.bUnion id), + { rw [sub_mul, one_mul, sub_le_sub_iff_left], + refine mul_le_of_le_one_right (by positivity) _, + exact_mod_cast G.edge_density_le_one _ _ }, + refine this.trans _, + simp only [simple_graph.edge_density_def, simple_graph.interedges, ←sup_eq_bUnion, cast_sum, + rel.card_interedges_finpartition _ (of_subset _ hA rfl) (of_subset _ hB rfl), of_subset_parts, + sum_div, mul_sum, rat.cast_sum, rat.cast_div, rat.cast_mul, div_div, + mul_div_left_comm ((1:ℝ) - _)], + push_cast, + apply sum_le_sum, + simp only [and_imp, prod.forall, mem_product], + rintro x y hx hy, + rw [mul_mul_mul_comm, mul_comm (x.card : ℝ), mul_comm (y.card : ℝ), le_div_iff, mul_assoc], + { refine mul_le_of_le_one_right (cast_nonneg _) _, + rw [div_mul_eq_mul_div, ←mul_assoc, mul_assoc], + refine div_le_one_of_le _ (by positivity), + refine (mul_le_mul_of_nonneg_right (one_sub_le_m_div_m_add_one_sq hPα hPε) _).trans _, + { exact_mod_cast (zero_le _) }, + rw [sq, mul_mul_mul_comm, mul_comm ((m:ℝ)/_), mul_comm ((m:ℝ)/_)], + refine mul_le_mul _ _ _ (cast_nonneg _), + apply le_sum_card_subset_chunk_parts hA hx, + apply le_sum_card_subset_chunk_parts hB hy, + positivity }, + refine mul_pos (mul_pos _ _) (mul_pos _ _); rw [cast_pos, finset.card_pos], + exacts [⟨_, hx⟩, nonempty_of_mem_parts _ (hA hx), ⟨_, hy⟩, nonempty_of_mem_parts _ (hB hy)] +end + +private lemma sum_density_div_card_le_density_add_eps [nonempty α] + (hPα : P.parts.card * 16 ^ P.parts.card ≤ card α) (hPε : 100 ≤ 4 ^ P.parts.card * ε ^ 5) + (hε₁ : ε ≤ 1) {hU : U ∈ P.parts} {hV : V ∈ P.parts} {A B : finset (finset α)} + (hA : A ⊆ (chunk hP G ε hU).parts) (hB : B ⊆ (chunk hP G ε hV).parts) : + (∑ ab in A.product B, G.edge_density ab.1 ab.2 : ℝ) / (A.card * B.card) ≤ + G.edge_density (A.bUnion id) (B.bUnion id) + ε ^ 5 / 49 := +begin + have : (1 + ε ^ 5/49) * G.edge_density (A.bUnion id) (B.bUnion id) ≤ + G.edge_density (A.bUnion id) (B.bUnion id) + ε ^ 5 / 49, + { rw [add_mul, one_mul, add_le_add_iff_left], + refine mul_le_of_le_one_right (by positivity) _, + exact_mod_cast G.edge_density_le_one _ _ }, + refine le_trans _ this, + simp only [simple_graph.edge_density, edge_density, ←sup_eq_bUnion, cast_sum, mul_sum, sum_div, + rel.card_interedges_finpartition _ (of_subset _ hA rfl) (of_subset _ hB rfl), rat.cast_sum, + rat.cast_div, rat.cast_mul, of_subset_parts, mul_div_left_comm ((1:ℝ) + _), div_div], + push_cast, + apply sum_le_sum, + simp only [and_imp, prod.forall, mem_product], + intros x y hx hy, + rw [mul_mul_mul_comm, mul_comm (x.card : ℝ), mul_comm (y.card : ℝ), div_le_iff, mul_assoc], + { refine le_mul_of_one_le_right (cast_nonneg _) _, + rw [div_mul_eq_mul_div, one_le_div], + refine le_trans _ (mul_le_mul_of_nonneg_right (m_add_one_div_m_le_one_add hPα hPε hε₁) _), + { rw [sq, mul_mul_mul_comm, mul_comm (_/(m:ℝ)), mul_comm (_/(m:ℝ))], + exact mul_le_mul (sum_card_subset_chunk_parts_le (by positivity) hA hx) + (sum_card_subset_chunk_parts_le (by positivity) hB hy) (by positivity) (by positivity) }, + { exact_mod_cast (zero_le _) }, + rw [←cast_mul, cast_pos], + apply mul_pos; rw [finset.card_pos, sup_eq_bUnion, bUnion_nonempty], + { exact ⟨_, hx, nonempty_of_mem_parts _ (hA hx)⟩ }, + { exact ⟨_, hy, nonempty_of_mem_parts _ (hB hy)⟩ } }, + refine mul_pos (mul_pos _ _) (mul_pos _ _); rw [cast_pos, finset.card_pos], + exacts [⟨_, hx⟩, nonempty_of_mem_parts _ (hA hx), ⟨_, hy⟩, nonempty_of_mem_parts _ (hB hy)] +end + +private lemma average_density_near_total_density [nonempty α] + (hPα : P.parts.card * 16 ^ P.parts.card ≤ card α) (hPε : 100 ≤ 4 ^ P.parts.card * ε ^ 5) + (hε₁ : ε ≤ 1) {hU : U ∈ P.parts} {hV : V ∈ P.parts} {A B : finset (finset α)} + (hA : A ⊆ (chunk hP G ε hU).parts) (hB : B ⊆ (chunk hP G ε hV).parts) : + |(∑ ab in A.product B, G.edge_density ab.1 ab.2 : ℝ)/(A.card * B.card) - + G.edge_density (A.bUnion id) (B.bUnion id)| ≤ ε ^ 5/49 := +begin + rw abs_sub_le_iff, + split, + { rw sub_le_iff_le_add', + exact sum_density_div_card_le_density_add_eps hPα hPε hε₁ hA hB }, + suffices : (G.edge_density (A.bUnion id) (B.bUnion id) : ℝ) - + (∑ ab in A.product B, G.edge_density ab.1 ab.2)/(A.card * B.card) ≤ ε ^ 5/50, + { apply this.trans, + exact div_le_div_of_le_left (by positivity) (by norm_num) (by norm_num) }, + rw [sub_le_iff_le_add, ←sub_le_iff_le_add'], + apply density_sub_eps_le_sum_density_div_card hPα hPε hA hB, +end + +private lemma edge_density_chunk_aux [nonempty α] + (hPα : P.parts.card * 16 ^ P.parts.card ≤ card α) (hPε : 100 ≤ 4 ^ P.parts.card * ε ^ 5) + (hU : U ∈ P.parts) (hV : V ∈ P.parts) : + ↑(G.edge_density U V) ^ 2 - ε ^ 5/25 ≤ + ((∑ ab in (chunk hP G ε hU).parts.product (chunk hP G ε hV).parts, + G.edge_density ab.1 ab.2)/16 ^ P.parts.card) ^ 2 := +begin + obtain hGε | hGε := le_total ↑(G.edge_density U V) (ε ^ 5/50), + { refine (sub_nonpos_of_le $ (sq_le _ _).trans $ hGε.trans _).trans (sq_nonneg _), + { exact_mod_cast G.edge_density_nonneg _ _ }, + { exact_mod_cast G.edge_density_le_one _ _ }, + { exact div_le_div_of_le_left (by positivity) (by norm_num) (by norm_num) } }, + rw ←sub_nonneg at hGε, + have : ↑(G.edge_density U V) - ε ^ 5 / 50 ≤ + (∑ ab in (chunk hP G ε hU).parts.product (chunk hP G ε hV).parts, + G.edge_density ab.1 ab.2) / (16 ^ P.parts.card), + { refine (le_trans _ $ density_sub_eps_le_sum_density_div_card hPα hPε + (set.subset.refl (chunk hP G ε hU).parts) + (set.subset.refl (chunk hP G ε hV).parts)).trans _, + { rw [bUnion_parts, bUnion_parts] }, + { rw [card_chunk (m_pos hPα).ne', card_chunk (m_pos hPα).ne', ←cast_mul, + ←mul_pow, cast_pow], + norm_cast } }, + refine le_trans _ (pow_le_pow_of_le_left hGε this 2), + rw [sub_sq, sub_add, sub_le_sub_iff_left], + refine (sub_le_self _ $ sq_nonneg $ ε ^ 5/50).trans _, + rw [mul_right_comm, mul_div_left_comm, div_eq_mul_inv (ε ^ 5), show (2:ℝ)/50 = 25⁻¹, by norm_num], + exact mul_le_of_le_one_right (by positivity) (by exact_mod_cast G.edge_density_le_one _ _), +end + +private lemma abs_density_star_sub_density_le_eps (hPε : 100 ≤ 4 ^ P.parts.card * ε ^ 5) + (hε₁ : ε ≤ 1) {hU : U ∈ P.parts} {hV : V ∈ P.parts} (hUV' : U ≠ V) (hUV : ¬ G.is_uniform ε U V) : + |(G.edge_density ((star hP G ε hU V).bUnion id) ((star hP G ε hV U).bUnion id) : ℝ) - + G.edge_density (G.nonuniform_witness ε U V) (G.nonuniform_witness ε V U)| ≤ ε/5 := +begin + convert abs_edge_density_sub_edge_density_le_two_mul G.adj + (bUnion_star_subset_nonuniform_witness hP G ε hU V) + (bUnion_star_subset_nonuniform_witness hP G ε hV U) + (by positivity) + (one_sub_eps_mul_card_nonuniform_witness_le_card_star hV hUV' hUV hPε hε₁) + (one_sub_eps_mul_card_nonuniform_witness_le_card_star hU hUV'.symm (λ hVU, hUV hVU.symm) + hPε hε₁), + linarith, +end + +private lemma eps_le_card_star_div [nonempty α] (hPα : P.parts.card * 16 ^ P.parts.card ≤ card α) + (hPε : 100 ≤ 4 ^ P.parts.card * ε ^ 5) (hε₁ : ε ≤ 1) (hU : U ∈ P.parts) + (hV : V ∈ P.parts) (hUV : U ≠ V) (hunif : ¬ G.is_uniform ε U V) : + 4/5 * ε ≤ (star hP G ε hU V).card / 4 ^ P.parts.card := +begin + have hm : (0 : ℝ) ≤ 1 - m⁻¹ := sub_nonneg_of_le (inv_le_one $ one_le_m_coe hPα), + have hε : 0 ≤ 1 - ε / 10 := + sub_nonneg_of_le (div_le_one_of_le (hε₁.trans $ by norm_num) $ by norm_num), + calc + 4/5 * ε + = (1 - 1/10) * (1 - 9⁻¹) * ε : by norm_num + ... ≤ (1 - ε/10) * (1 - m⁻¹) * ((G.nonuniform_witness ε U V).card / U.card) + : mul_le_mul + (mul_le_mul (sub_le_sub_left (div_le_div_of_le_of_nonneg hε₁ $ by norm_num) _) + (sub_le_sub_left (inv_le_inv_of_le (by norm_num) $ + by exact_mod_cast (show 9 ≤ 100, by norm_num).trans (hundred_le_m hPα hPε hε₁)) _) + (by norm_num) hε) + ((le_div_iff' $ (@cast_pos ℝ _ _ _).2 (P.nonempty_of_mem_parts hU).card_pos).2 $ + G.le_card_nonuniform_witness hunif) + (by positivity) (by positivity) + ... = (1 - ε/10) * (G.nonuniform_witness ε U V).card * ((1 - m⁻¹) / U.card) + : by rw [mul_assoc, mul_assoc, mul_div_left_comm] + ... ≤ ((star hP G ε hU V).bUnion id).card * ((1 - m⁻¹) / U.card) + : mul_le_mul_of_nonneg_right + (one_sub_eps_mul_card_nonuniform_witness_le_card_star hV hUV hunif hPε hε₁) + (by positivity) + ... ≤ (star hP G ε hU V).card * (m + 1) * ((1 - m⁻¹) / U.card) : + mul_le_mul_of_nonneg_right card_bUnion_star_le_m_add_one_card_star_mul (by positivity) + ... ≤ (star hP G ε hU V).card * (m + 1) * ((1 - m⁻¹) / (4 ^ P.parts.card * m)) + : mul_le_mul_of_nonneg_left (div_le_div_of_le_left hm (by positivity) $ + pow_mul_m_le_card_part hP hU) (by positivity) + ... ≤ (star hP G ε hU V).card / 4 ^ P.parts.card : + begin + rw [mul_assoc, mul_comm ((4:ℝ)^P.parts.card), ←div_div, ←mul_div_assoc, ←mul_comm_div], + refine mul_le_of_le_one_right (by positivity) _, + have hm : (0 : ℝ) < m := by positivity, + rw [mul_div_assoc', div_le_one hm, ←one_div, one_sub_div hm.ne', mul_div_assoc', + div_le_iff hm], + linarith, + end +end + +/-! +### Final bounds + +Those inequalities are the end result of all this hard work. +-/ + +/-- Lower bound on the edge densities between non-uniform parts of `szemeredi_regularity.star`. -/ +private lemma edge_density_star_not_uniform [nonempty α] + (hPα : P.parts.card * 16 ^ P.parts.card ≤ card α) (hPε : 100 ≤ 4 ^ P.parts.card * ε ^ 5) + (hε₁ : ε ≤ 1) {hU : U ∈ P.parts} {hV : V ∈ P.parts} (hUVne : U ≠ V) (hUV : ¬ G.is_uniform ε U V) : + 3/4 * ε ≤ + |(∑ ab in (star hP G ε hU V).product (star hP G ε hV U), G.edge_density ab.1 ab.2) + / ((star hP G ε hU V).card * (star hP G ε hV U).card) - + (∑ ab in (chunk hP G ε hU).parts.product (chunk hP G ε hV).parts, + G.edge_density ab.1 ab.2)/16 ^ P.parts.card| := +begin + rw [(show (16:ℝ) = 4 ^ 2, by norm_num), pow_right_comm, sq ((4:ℝ)^_)], + set p : ℝ := (∑ ab in (star hP G ε hU V).product (star hP G ε hV U), G.edge_density ab.1 ab.2) + / ((star hP G ε hU V).card * (star hP G ε hV U).card), + set q : ℝ := (∑ ab in (chunk hP G ε hU).parts.product (chunk hP G ε hV).parts, + G.edge_density ab.1 ab.2)/(4 ^ P.parts.card * 4 ^ P.parts.card), + change _ ≤ |p - q|, + set r : ℝ := G.edge_density ((star hP G ε hU V).bUnion id) ((star hP G ε hV U).bUnion id), + set s : ℝ := G.edge_density (G.nonuniform_witness ε U V) (G.nonuniform_witness ε V U), + set t : ℝ := G.edge_density U V, + have hrs : |r - s| ≤ ε/5 := abs_density_star_sub_density_le_eps hPε hε₁ hUVne hUV, + have hst : ε ≤ |s - t| := G.nonuniform_witness_spec hUVne hUV, + have hpr : |p - r| ≤ ε ^ 5/49 := average_density_near_total_density hPα hPε hε₁ + star_subset_chunk star_subset_chunk, + have hqt : |q - t| ≤ ε ^ 5/49, + { have := average_density_near_total_density hPα hPε hε₁ + (subset.refl (chunk hP G ε hU).parts) + (subset.refl (chunk hP G ε hV).parts), + simp_rw [←sup_eq_bUnion, sup_parts, card_chunk (m_pos hPα).ne', cast_pow] at this, + norm_num at this, + exact this }, + have hε' : ε ^ 5 ≤ ε := by simpa using pow_le_pow_of_le_one (by positivity) hε₁ + (show 1 ≤ 5, by norm_num), + have hpr' : |p - r| ≤ ε/49 := hpr.trans (div_le_div_of_le_of_nonneg hε' $ by norm_num), + have hqt' : |q - t| ≤ ε/49 := hqt.trans (div_le_div_of_le_of_nonneg hε' $ by norm_num), + rw abs_sub_le_iff at hrs hpr' hqt', + rw le_abs at hst ⊢, + cases hst, + left, linarith, + right, linarith, +end + +/-- Lower bound on the edge densities between non-uniform parts of `szemeredi_regularity.increment`. +-/ +lemma edge_density_chunk_not_uniform [nonempty α] + (hPα : P.parts.card * 16 ^ P.parts.card ≤ card α) (hPε : 100 ≤ 4 ^ P.parts.card * ε ^ 5) + (hε₁ : ε ≤ 1) {hU : U ∈ P.parts} {hV : V ∈ P.parts} (hUVne : U ≠ V) (hUV : ¬ G.is_uniform ε U V) : + (G.edge_density U V : ℝ) ^ 2 - ε ^ 5/25 + ε^4/3 ≤ + (∑ ab in (chunk hP G ε hU).parts.product (chunk hP G ε hV).parts, + G.edge_density ab.1 ab.2 ^ 2)/16 ^ P.parts.card := +calc + ↑(G.edge_density U V) ^ 2 - ε ^ 5/25 + ε^4/3 + ≤ G.edge_density U V ^ 2 - ε ^ 5/25 + + (star hP G ε hU V).card * (star hP G ε hV U).card/16 ^ P.parts.card * (9/16) * ε ^ 2 : + begin + apply add_le_add_left, + have Ul : 4/5 * ε ≤ (star hP G ε hU V).card / _ := + eps_le_card_star_div hPα hPε hε₁ hU hV hUVne hUV, + have Vl : 4/5 * ε ≤ (star hP G ε hV U).card / _ := + eps_le_card_star_div hPα hPε hε₁ hV hU hUVne.symm (λ h, hUV h.symm), + rw [(show (16 : ℝ) = 4 ^ 2, by norm_num), pow_right_comm, sq ((4:ℝ)^_), + ←_root_.div_mul_div_comm, mul_assoc], + have : 0 < ε := by positivity, + have UVl := mul_le_mul Ul Vl (by positivity) (by positivity), + refine le_trans _ (mul_le_mul_of_nonneg_right UVl _), + { field_simp, + ring_nf, + apply mul_le_mul_of_nonneg_right, + norm_num, + positivity }, + { norm_num, + positivity } + end + ... ≤ (∑ ab in (chunk hP G ε hU).parts.product (chunk hP G ε hV).parts, + G.edge_density ab.1 ab.2 ^ 2)/16 ^ P.parts.card : + begin + have t : (star hP G ε hU V).product (star hP G ε hV U) ⊆ + (chunk hP G ε hU).parts.product (chunk hP G ε hV).parts := + product_subset_product star_subset_chunk star_subset_chunk, + have hε : 0 ≤ ε := by positivity, + have := add_div_le_sum_sq_div_card t (λ x, (G.edge_density x.1 x.2 : ℝ)) + (G.edge_density U V ^ 2 - ε ^ 5/25) (show 0 ≤ 3/4 * ε, by linarith) _ _, + { simp_rw [card_product, card_chunk (m_pos hPα).ne', ←mul_pow, cast_pow, mul_pow, div_pow, + ←mul_assoc] at this, + norm_num at this, + exact this }, + { simp_rw [card_product, card_chunk (m_pos hPα).ne', ←mul_pow], + norm_num, + exact edge_density_star_not_uniform hPα hPε hε₁ hUVne hUV }, + { rw card_product, + apply (edge_density_chunk_aux hPα hPε hU hV).trans, + rw [card_chunk (m_pos hPα).ne', card_chunk (m_pos hPα).ne', ←mul_pow], + norm_num, + exact hP } + end + +/-- Lower bound on the edge densities between parts of `szemeredi_regularity.increment`. This is the +blanket lower bound used the uniform parts. -/ +lemma edge_density_chunk_uniform [nonempty α] + (hPα : P.parts.card * 16 ^ P.parts.card ≤ card α) (hPε : 100 ≤ 4 ^ P.parts.card * ε ^ 5) + (hU : U ∈ P.parts) (hV : V ∈ P.parts) : + (G.edge_density U V : ℝ) ^ 2 - ε ^ 5/25 ≤ + (∑ ab in (chunk hP G ε hU).parts.product (chunk hP G ε hV).parts, + G.edge_density ab.1 ab.2 ^ 2)/16 ^ P.parts.card := +begin + apply (edge_density_chunk_aux hPα hPε hU hV).trans, + convert sum_div_card_sq_le_sum_sq_div_card; + rw [card_product, cast_mul, card_chunk (m_pos hPα).ne', card_chunk (m_pos hPα).ne', ←cast_mul, + ←mul_pow]; + norm_cast, +end + +end szemeredi_regularity diff --git a/src/combinatorics/simple_graph/regularity/uniform.lean b/src/combinatorics/simple_graph/regularity/uniform.lean index 47a34c7cc18de..57f1c1ecd6df1 100644 --- a/src/combinatorics/simple_graph/regularity/uniform.lean +++ b/src/combinatorics/simple_graph/regularity/uniform.lean @@ -146,7 +146,7 @@ begin { exact G.right_nonuniform_witnesses_subset (λ i, h i.symm) } end -lemma nonuniform_witness_card_le (h : ¬ G.is_uniform ε s t) : +lemma le_card_nonuniform_witness (h : ¬ G.is_uniform ε s t) : (s.card : 𝕜) * ε ≤ (G.nonuniform_witness ε s t).card := begin unfold nonuniform_witness, From bf7ef0e83e5b7e6c1169e97f055e58a2e4e9d52d Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 19 May 2023 22:52:50 +0000 Subject: [PATCH 1043/1291] feat(combinatorics/simple_graph/regularity): Increment partition (#19051) Define the increment partition and prove its two crucial properties: * It has size depending only on the size of the original partition * It increases the energy by a fixed amount This is all internal to the proof of SRL, so I made most lemmas `private`. Co-authored-by: Bhavik Mehta --- docs/references.bib | 23 +++ .../simple_graph/regularity/bound.lean | 4 + .../simple_graph/regularity/chunk.lean | 10 +- .../simple_graph/regularity/energy.lean | 9 + .../simple_graph/regularity/equitabilise.lean | 4 + .../simple_graph/regularity/increment.lean | 195 ++++++++++++++++++ .../simple_graph/regularity/uniform.lean | 4 + 7 files changed, 246 insertions(+), 3 deletions(-) create mode 100644 src/combinatorics/simple_graph/regularity/increment.lean diff --git a/docs/references.bib b/docs/references.bib index d57e87b338103..470569c9a9037 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -2057,6 +2057,29 @@ @Book{ soare1987 doi = {10.1007/978-3-662-02460-7} } +@InProceedings{ srl_itp, + author = {Dillies, Ya\"{e}l and Mehta, Bhavik}, + title = {{Formalising Szemer\'{e}di’s Regularity Lemma in Lean}}, + booktitle = {13th International Conference on Interactive Theorem + Proving (ITP 2022)}, + pages = {9:1--9:19}, + series = {Leibniz International Proceedings in Informatics + (LIPIcs)}, + isbn = {978-3-95977-252-5}, + issn = {1868-8969}, + year = {2022}, + volume = {237}, + editor = {Andronick, June and de Moura, Leonardo}, + publisher = {Schloss Dagstuhl -- Leibniz-Zentrum f{\"u}r Informatik}, + address = {Dagstuhl, Germany}, + url = {https://drops.dagstuhl.de/opus/volltexte/2022/16718}, + urn = {urn:nbn:de:0030-drops-167185}, + doi = {10.4230/LIPIcs.ITP.2022.9}, + annote = {Keywords: Lean, formalisation, formal proof, graph theory, + combinatorics, additive combinatorics, Szemer\'{e}di’s + Regularity Lemma, Roth’s Theorem} +} + @Book{ stanley2012, author = {Stanley, Richard P.}, title = {Enumerative combinatorics}, diff --git a/src/combinatorics/simple_graph/regularity/bound.lean b/src/combinatorics/simple_graph/regularity/bound.lean index 41e723527b1c1..2241cf244ebaf 100644 --- a/src/combinatorics/simple_graph/regularity/bound.lean +++ b/src/combinatorics/simple_graph/regularity/bound.lean @@ -21,6 +21,10 @@ This entire file is internal to the proof of Szemerédi Regularity Lemma. * `szemeredi_regularity.initial_bound`: The size of the partition we start the induction with. * `szemeredi_regularity.bound`: The upper bound on the size of the partition produced by our version of Szemerédi's regularity lemma. + +## References + +[Yaël Dillies, Bhavik Mehta, *Formalising Szemerédi’s Regularity Lemma in Lean*][srl_itp] -/ open finset fintype function real diff --git a/src/combinatorics/simple_graph/regularity/chunk.lean b/src/combinatorics/simple_graph/regularity/chunk.lean index e72a5221269bf..ab766c66b386c 100644 --- a/src/combinatorics/simple_graph/regularity/chunk.lean +++ b/src/combinatorics/simple_graph/regularity/chunk.lean @@ -28,6 +28,10 @@ This entire file is internal to the proof of Szemerédi Regularity Lemma. Once ported to mathlib4, this file will be a great golfing ground for Heather's new tactic `rel_congr`. + +## References + +[Yaël Dillies, Bhavik Mehta, *Formalising Szemerédi’s Regularity Lemma in Lean*][srl_itp] -/ open finpartition finset fintype rel nat @@ -50,9 +54,9 @@ contained in the corresponding witness of non-uniformity. /-- The portion of `szemeredi_regularity.increment` which partitions `U`. -/ noncomputable def chunk : finpartition U := -dite (U.card = m * 4 ^ P.parts.card + (card α/P.parts.card - m * 4 ^ P.parts.card)) - (λ hUcard, (atomise U $ P.nonuniform_witnesses G ε U).equitabilise $ card_aux₁ hUcard) - (λ hUcard, (atomise U $ P.nonuniform_witnesses G ε U).equitabilise $ card_aux₂ hP hU hUcard) +if hUcard : U.card = m * 4 ^ P.parts.card + (card α/P.parts.card - m * 4 ^ P.parts.card) + then (atomise U $ P.nonuniform_witnesses G ε U).equitabilise $ card_aux₁ hUcard + else (atomise U $ P.nonuniform_witnesses G ε U).equitabilise $ card_aux₂ hP hU hUcard -- `hP` and `hU` are used to get that `U` has size -- `m * 4 ^ P.parts.card + a or m * 4 ^ P.parts.card + a + 1` diff --git a/src/combinatorics/simple_graph/regularity/energy.lean b/src/combinatorics/simple_graph/regularity/energy.lean index a4624af98ae20..e4e91107fb061 100644 --- a/src/combinatorics/simple_graph/regularity/energy.lean +++ b/src/combinatorics/simple_graph/regularity/energy.lean @@ -6,6 +6,7 @@ Authors: Yaël Dillies, Bhavik Mehta import algebra.big_operators.order import algebra.module.basic import combinatorics.simple_graph.density +import data.rat.big_operators /-! # Energy of a partition @@ -18,6 +19,10 @@ This file defines the energy of a partition. The energy is the auxiliary quantity that drives the induction process in the proof of Szemerédi's Regularity Lemma. As long as we do not have a suitable equipartition, we will find a new one that has an energy greater than the previous one plus some fixed constant. + +## References + +[Yaël Dillies, Bhavik Mehta, *Formalising Szemerédi’s Regularity Lemma in Lean*][srl_itp] -/ open finset @@ -44,4 +49,8 @@ div_le_of_nonneg_of_le_mul (sq_nonneg _) zero_le_one $ ... = P.parts.off_diag.card : nat.smul_one_eq_coe _ ... ≤ _ : by { rw [off_diag_card, one_mul, ←nat.cast_pow, nat.cast_le, sq], exact tsub_le_self } +@[simp, norm_cast] lemma coe_energy {𝕜 : Type*} [linear_ordered_field 𝕜] : + (P.energy G : 𝕜) = (∑ uv in P.parts.off_diag, G.edge_density uv.1 uv.2 ^ 2) / P.parts.card ^ 2 := +by { rw energy, norm_cast } + end finpartition diff --git a/src/combinatorics/simple_graph/regularity/equitabilise.lean b/src/combinatorics/simple_graph/regularity/equitabilise.lean index 32e964da201a5..35531993e5767 100644 --- a/src/combinatorics/simple_graph/regularity/equitabilise.lean +++ b/src/combinatorics/simple_graph/regularity/equitabilise.lean @@ -22,6 +22,10 @@ This file allows to blow partitions up into parts of controlled size. Given a pa * `finpartition.equitabilise`: `P.equitabilise h` where `h : a * m + b * (m + 1)` is a partition with `a` parts of size `m` and `b` parts of size `m + 1` which almost refines `P`. * `finpartition.exists_equipartition_card_eq`: We can find equipartitions of arbitrary size. + +## References + +[Yaël Dillies, Bhavik Mehta, *Formalising Szemerédi’s Regularity Lemma in Lean*][srl_itp] -/ open finset nat diff --git a/src/combinatorics/simple_graph/regularity/increment.lean b/src/combinatorics/simple_graph/regularity/increment.lean new file mode 100644 index 0000000000000..8f0c29a8e2a28 --- /dev/null +++ b/src/combinatorics/simple_graph/regularity/increment.lean @@ -0,0 +1,195 @@ +/- +Copyright (c) 2022 Yaël Dillies, Bhavik Mehta. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies, Bhavik Mehta +-/ +import combinatorics.simple_graph.regularity.chunk +import combinatorics.simple_graph.regularity.energy + +/-! +# Increment partition for Szemerédi Regularity Lemma + +In the proof of Szemerédi Regularity Lemma, we need to partition each part of a starting partition +to increase the energy. This file defines the partition obtained by gluing the parts partitions +together (the *increment partition*) and shows that the energy globally increases. + +This entire file is internal to the proof of Szemerédi Regularity Lemma. + +## Main declarations + +* `szemeredi_regularity.increment`: The increment partition. +* `szemeredi_regularity.card_increment`: The increment partition is much bigger than the original, + but by a controlled amount. +* `szemeredi_regularity.energy_increment`: The increment partition has energy greater than the + original by a known (small) fixed amount. + +## TODO + +Once ported to mathlib4, this file will be a great golfing ground for Heather's new tactic +`rel_congr`. + +## References + +[Yaël Dillies, Bhavik Mehta, *Formalising Szemerédi’s Regularity Lemma in Lean*][srl_itp] +-/ + +open finset fintype simple_graph szemeredi_regularity +open_locale big_operators classical + +local attribute [positivity] tactic.positivity_szemeredi_regularity + +variables {α : Type*} [fintype α] {P : finpartition (univ : finset α)} (hP : P.is_equipartition) + (G : simple_graph α) (ε : ℝ) + +local notation `m` := (card α/step_bound P.parts.card : ℕ) + +namespace szemeredi_regularity + +/-- The **increment partition** in Szemerédi's Regularity Lemma. + +If an equipartition is *not* uniform, then the increment partition is a (much bigger) equipartition +with a slightly higher energy. This is helpful since the energy is bounded by a constant (see +`szemeredi_regularity.energy_le_one`), so this process eventually terminates and yields a +not-too-big uniform equipartition. -/ +noncomputable def increment : finpartition (univ : finset α) := P.bind $ λ U, chunk hP G ε + +open finpartition finpartition.is_equipartition + +variables {hP G ε} + +/-- The increment partition has a prescribed (very big) size in terms of the original partition. -/ +lemma card_increment (hPα : P.parts.card * 16^P.parts.card ≤ card α) (hPG : ¬P.is_uniform G ε) : + (increment hP G ε).parts.card = step_bound P.parts.card := +begin + have hPα' : step_bound P.parts.card ≤ card α := + (mul_le_mul_left' (pow_le_pow_of_le_left' (by norm_num) _) _).trans hPα, + have hPpos : 0 < step_bound P.parts.card := step_bound_pos (nonempty_of_not_uniform hPG).card_pos, + rw [increment, card_bind], + simp_rw [chunk, apply_dite finpartition.parts, apply_dite card, sum_dite], + rw [sum_const_nat, sum_const_nat, card_attach, card_attach], rotate, + any_goals { exact λ x hx, card_parts_equitabilise _ _ (nat.div_pos hPα' hPpos).ne' }, + rw [nat.sub_add_cancel a_add_one_le_four_pow_parts_card, nat.sub_add_cancel ((nat.le_succ _).trans + a_add_one_le_four_pow_parts_card), ←add_mul], + congr, + rw [filter_card_add_filter_neg_card_eq_card, card_attach], +end + +lemma increment_is_equipartition (hP : P.is_equipartition) (G : simple_graph α) (ε : ℝ) : + (increment hP G ε).is_equipartition := +begin + simp_rw [is_equipartition, set.equitable_on_iff_exists_eq_eq_add_one], + refine ⟨m, λ A hA, _⟩, + rw [mem_coe, increment, mem_bind] at hA, + obtain ⟨U, hU, hA⟩ := hA, + exact card_eq_of_mem_parts_chunk hA, +end + +private lemma distinct_pairs_increment : + P.parts.off_diag.attach.bUnion + (λ UV, (chunk hP G ε (mem_off_diag.1 UV.2).1).parts ×ˢ + (chunk hP G ε (mem_off_diag.1 UV.2).2.1).parts) + ⊆ (increment hP G ε).parts.off_diag := +begin + rintro ⟨Ui, Vj⟩, + simp only [increment, mem_off_diag, bind_parts, mem_bUnion, prod.exists, exists_and_distrib_left, + exists_prop, mem_product, mem_attach, true_and, subtype.exists, and_imp, mem_off_diag, + forall_exists_index, bex_imp_distrib, ne.def], + refine λ U V hUV hUi hVj, ⟨⟨_, hUV.1, hUi⟩, ⟨_, hUV.2.1, hVj⟩, _⟩, + rintro rfl, + obtain ⟨i, hi⟩ := nonempty_of_mem_parts _ hUi, + exact hUV.2.2 (P.disjoint.elim_finset hUV.1 hUV.2.1 i (finpartition.le _ hUi hi) $ + finpartition.le _ hVj hi), +end + +/-- The contribution to `finpartition.energy` of a pair of distinct parts of a finpartition. -/ +private noncomputable def pair_contrib (G : simple_graph α) (ε : ℝ) (hP : P.is_equipartition) + (x : {x // x ∈ P.parts.off_diag}) : ℚ := +∑ i in (chunk hP G ε (mem_off_diag.1 x.2).1).parts ×ˢ (chunk hP G ε (mem_off_diag.1 x.2).2.1).parts, + G.edge_density i.fst i.snd ^ 2 + +lemma off_diag_pairs_le_increment_energy : + ∑ x in P.parts.off_diag.attach, pair_contrib G ε hP x / (increment hP G ε).parts.card ^ 2 ≤ + (increment hP G ε).energy G := +begin + simp_rw [pair_contrib, ←sum_div], + refine div_le_div_of_le_of_nonneg _ (sq_nonneg _), + rw ←sum_bUnion, + { exact sum_le_sum_of_subset_of_nonneg distinct_pairs_increment (λ i _ _, sq_nonneg _) }, + simp only [set.pairwise_disjoint, function.on_fun, disjoint_left, inf_eq_inter, mem_inter, + mem_product], + rintro ⟨⟨s₁, s₂⟩, hs⟩ _ ⟨⟨t₁, t₂⟩, ht⟩ _ hst ⟨u, v⟩ huv₁ huv₂, + rw mem_off_diag at hs ht, + obtain ⟨a, ha⟩ := finpartition.nonempty_of_mem_parts _ huv₁.1, + obtain ⟨b, hb⟩ := finpartition.nonempty_of_mem_parts _ huv₁.2, + exact hst (subtype.ext_val $ prod.ext + (P.disjoint.elim_finset hs.1 ht.1 a + (finpartition.le _ huv₁.1 ha) $ finpartition.le _ huv₂.1 ha) $ + P.disjoint.elim_finset hs.2.1 ht.2.1 b + (finpartition.le _ huv₁.2 hb) $ finpartition.le _ huv₂.2 hb), +end + +lemma pair_contrib_lower_bound [nonempty α] (x : {i // i ∈ P.parts.off_diag}) (hε₁ : ε ≤ 1) + (hPα : P.parts.card * 16^P.parts.card ≤ card α) (hPε : 100 ≤ 4^P.parts.card * ε^5) : + ↑(G.edge_density x.1.1 x.1.2)^2 - ε^5/25 + (if G.is_uniform ε x.1.1 x.1.2 then 0 else ε^4/3) ≤ + pair_contrib G ε hP x / (16^P.parts.card) := +begin + rw pair_contrib, + push_cast, + split_ifs, + { rw add_zero, + exact edge_density_chunk_uniform hPα hPε _ _ }, + { exact edge_density_chunk_not_uniform hPα hPε hε₁ (mem_off_diag.1 x.2).2.2 h } +end + +lemma uniform_add_nonuniform_eq_off_diag_pairs [nonempty α] (hε₁ : ε ≤ 1) (hP₇ : 7 ≤ P.parts.card) + (hPα : P.parts.card * 16^P.parts.card ≤ card α) (hPε : 100 ≤ 4^P.parts.card * ε^5) + (hPG : ¬P.is_uniform G ε) : + (∑ x in P.parts.off_diag, G.edge_density x.1 x.2 ^ 2 + P.parts.card^2 * (ε ^ 5 / 4) : ℝ) + / P.parts.card ^ 2 + ≤ ∑ x in P.parts.off_diag.attach, pair_contrib G ε hP x / (increment hP G ε).parts.card ^ 2 := +begin + conv_rhs + { rw [←sum_div, card_increment hPα hPG, step_bound, ←nat.cast_pow, mul_pow, pow_right_comm, + nat.cast_mul, mul_comm, ←div_div, (show 4^2 = 16, by norm_num), sum_div] }, + rw [←nat.cast_pow, nat.cast_pow 16], + refine div_le_div_of_le_of_nonneg _ (nat.cast_nonneg _), + norm_num, + transitivity ∑ x in P.parts.off_diag.attach, + (G.edge_density x.1.1 x.1.2^2 - ε^5/25 + if G.is_uniform ε x.1.1 x.1.2 then 0 else ε^4/3 : ℝ), + swap, + { exact sum_le_sum (λ i hi, pair_contrib_lower_bound i hε₁ hPα hPε) }, + have : ∑ x in P.parts.off_diag.attach, + (G.edge_density x.1.1 x.1.2^2 - ε^5/25 + if G.is_uniform ε x.1.1 x.1.2 then 0 else ε^4/3 : ℝ) = + ∑ x in P.parts.off_diag, + (G.edge_density x.1 x.2^2 - ε^5/25 + if G.is_uniform ε x.1 x.2 then 0 else ε^4/3), + { convert sum_attach, refl }, + rw [this, sum_add_distrib, sum_sub_distrib, sum_const, nsmul_eq_mul, sum_ite, sum_const_zero, + zero_add, sum_const, nsmul_eq_mul, ←finpartition.non_uniforms], + rw [finpartition.is_uniform, not_le] at hPG, + refine le_trans _ (add_le_add_left (mul_le_mul_of_nonneg_right hPG.le $ by positivity) _), + conv_rhs { congr, congr, skip, rw [off_diag_card], congr, congr, + conv { congr, skip, rw ←mul_one P.parts.card }, rw ←nat.mul_sub_left_distrib }, + simp_rw [mul_assoc, sub_add_eq_add_sub, add_sub_assoc, ←mul_sub_left_distrib, mul_div_assoc' ε, + ←pow_succ, div_eq_mul_one_div (ε^5), ←mul_sub_left_distrib, mul_left_comm _ (ε^5), sq, + nat.cast_mul, mul_assoc, ←mul_assoc (ε ^ 5)], + refine add_le_add_left (mul_le_mul_of_nonneg_left _ $ by positivity) _, + rw [nat.cast_sub (P.parts_nonempty $ univ_nonempty.ne_empty).card_pos, mul_sub_right_distrib, + nat.cast_one, one_mul, le_sub_comm, ←mul_sub_left_distrib, + ←div_le_iff (show (0:ℝ) < 1/3 - 1/25 - 1/4, by norm_num)], + exact le_trans (show _ ≤ (7:ℝ), by norm_num) (by exact_mod_cast hP₇), +end + +/-- The increment partition has energy greater than the original one by a known fixed amount. -/ +lemma energy_increment [nonempty α] (hP : P.is_equipartition) (hP₇ : 7 ≤ P.parts.card) + (hε : 100 < 4^P.parts.card * ε^5) (hPα : P.parts.card * 16^P.parts.card ≤ card α) + (hPG : ¬P.is_uniform G ε) (hε₁ : ε ≤ 1) : + ↑(P.energy G) + ε^5 / 4 ≤ (increment hP G ε).energy G := +begin + rw coe_energy, + have h := uniform_add_nonuniform_eq_off_diag_pairs hε₁ hP₇ hPα hε.le hPG, + rw [add_div, mul_div_cancel_left] at h, + exact h.trans (by exact_mod_cast off_diag_pairs_le_increment_energy), + positivity, +end + +end szemeredi_regularity diff --git a/src/combinatorics/simple_graph/regularity/uniform.lean b/src/combinatorics/simple_graph/regularity/uniform.lean index 57f1c1ecd6df1..226cc63112a9a 100644 --- a/src/combinatorics/simple_graph/regularity/uniform.lean +++ b/src/combinatorics/simple_graph/regularity/uniform.lean @@ -33,6 +33,10 @@ is less than `ε`. * `finpartition.is_uniform`: Uniformity of a partition. * `finpartition.nonuniform_witnesses`: For each non-uniform pair of parts of a partition, pick witnesses of non-uniformity and dump them all together. + +## References + +[Yaël Dillies, Bhavik Mehta, *Formalising Szemerédi’s Regularity Lemma in Lean*][srl_itp] -/ open finset From 9b33e5f30c5f161e1d1b16b6b9b922bf49943377 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sat, 20 May 2023 10:41:26 +0000 Subject: [PATCH 1044/1291] feat(measure_theory/function/l1_space): generalize from fields to rings (#19052) --- src/measure_theory/function/l1_space.lean | 66 ++++++++++++++++------- 1 file changed, 48 insertions(+), 18 deletions(-) diff --git a/src/measure_theory/function/l1_space.lean b/src/measure_theory/function/l1_space.lean index 89b1e64eeb428..5a403b13399f0 100644 --- a/src/measure_theory/function/l1_space.lean +++ b/src/measure_theory/function/l1_space.lean @@ -350,15 +350,19 @@ hf.mono $ eventually_of_forall $ λ x, end pos_part section normed_space -variables {𝕜 : Type*} [normed_field 𝕜] [normed_space 𝕜 β] +variables {𝕜 : Type*} -lemma has_finite_integral.smul (c : 𝕜) {f : α → β} : has_finite_integral f μ → - has_finite_integral (c • f) μ := +lemma has_finite_integral.smul' [has_smul 𝕜 β] [has_nnnorm 𝕜] (c : 𝕜) {f : α → β} + (h : ∀ (k : 𝕜) (b : β), ‖k • b‖₊ ≤ ‖k‖₊ * ‖b‖₊) : + has_finite_integral f μ → has_finite_integral (c • f) μ := begin simp only [has_finite_integral], assume hfi, calc - ∫⁻ (a : α), ‖c • f a‖₊ ∂μ = ∫⁻ (a : α), (‖c‖₊) * ‖f a‖₊ ∂μ : - by simp only [nnnorm_smul, ennreal.coe_mul] + ∫⁻ (a : α), ‖c • f a‖₊ ∂μ ≤ ∫⁻ (a : α), (‖c‖₊) * ‖f a‖₊ ∂μ : begin + refine lintegral_mono _, + intro i, + exact_mod_cast h c (f i), + end ... < ∞ : begin rw lintegral_const_mul', @@ -366,7 +370,12 @@ begin end end -lemma has_finite_integral_smul_iff {c : 𝕜} (hc : c ≠ 0) (f : α → β) : +lemma has_finite_integral.smul [normed_field 𝕜] [normed_space 𝕜 β] (c : 𝕜) {f : α → β} : + has_finite_integral f μ → has_finite_integral (c • f) μ := +has_finite_integral.smul' _ $ λ a b, nnnorm_smul_le a b + +lemma has_finite_integral_smul_iff [normed_field 𝕜] [normed_space 𝕜 β] {c : 𝕜} (hc : c ≠ 0) + (f : α → β) : has_finite_integral (c • f) μ ↔ has_finite_integral f μ := begin split, @@ -375,13 +384,16 @@ begin exact has_finite_integral.smul _ end -lemma has_finite_integral.const_mul {f : α → ℝ} (h : has_finite_integral f μ) (c : ℝ) : +lemma has_finite_integral.const_mul [normed_ring 𝕜] {f : α → 𝕜} (h : has_finite_integral f μ) + (c : 𝕜) : has_finite_integral (λ x, c * f x) μ := -(has_finite_integral.smul c h : _) +(has_finite_integral.smul' c nnnorm_mul_le h : _) -lemma has_finite_integral.mul_const {f : α → ℝ} (h : has_finite_integral f μ) (c : ℝ) : +lemma has_finite_integral.mul_const [normed_ring 𝕜] {f : α → 𝕜} (h : has_finite_integral f μ) + (c : 𝕜) : has_finite_integral (λ x, f x * c) μ := -by simp_rw [mul_comm, h.const_mul _] +(has_finite_integral.smul' (mul_opposite.op c) + (λ a b, (nnnorm_mul_le b a.unop).trans_eq $ mul_comm _ _) h : _) end normed_space @@ -963,28 +975,32 @@ end end normed_space_over_complete_field -section is_R_or_C -variables {𝕜 : Type*} [is_R_or_C 𝕜] {f : α → 𝕜} +section normed_ring +variables {𝕜 : Type*} [normed_ring 𝕜] {f : α → 𝕜} lemma integrable.const_mul {f : α → 𝕜} (h : integrable f μ) (c : 𝕜) : integrable (λ x, c * f x) μ := -integrable.smul c h +⟨h.ae_strongly_measurable.const_smul c, h.has_finite_integral.const_mul c⟩ lemma integrable.const_mul' {f : α → 𝕜} (h : integrable f μ) (c : 𝕜) : integrable ((λ (x : α), c) * f) μ := -integrable.smul c h +integrable.const_mul h c lemma integrable.mul_const {f : α → 𝕜} (h : integrable f μ) (c : 𝕜) : integrable (λ x, f x * c) μ := -by simp_rw [mul_comm, h.const_mul _] +⟨h.ae_strongly_measurable.const_smul (mul_opposite.op c), h.has_finite_integral.mul_const c⟩ 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] +lemma integrable_const_mul_iff {c : 𝕜} (hc : is_unit c) (f : α → 𝕜) : + integrable (λ x, c * f x) μ ↔ integrable f μ := +let ⟨u, hc⟩ := hc in hc ▸ ⟨λ h, by simpa using h.const_mul ↑(u⁻¹), λ h, h.const_mul _⟩ + +lemma integrable_mul_const_iff {c : 𝕜} (hc : is_unit c) (f : α → 𝕜) : + integrable (λ x, f x * c) μ ↔ integrable f μ := +let ⟨u, hc⟩ := hc in hc ▸ ⟨λ h, by simpa using h.mul_const ↑(u⁻¹), λ h, h.mul_const _⟩ lemma integrable.bdd_mul' {f g : α → 𝕜} {c : ℝ} (hg : integrable g μ) (hf : ae_strongly_measurable f μ) (hf_bound : ∀ᵐ x ∂μ, ‖f x‖ ≤ c) : @@ -996,6 +1012,20 @@ begin exact (norm_mul_le _ _).trans (mul_le_mul_of_nonneg_right hx (norm_nonneg _)), end +end normed_ring + +section normed_division_ring +variables {𝕜 : Type*} [normed_division_ring 𝕜] {f : α → 𝕜} + +lemma integrable.div_const {f : α → 𝕜} (h : integrable f μ) (c : 𝕜) : + integrable (λ x, f x / c) μ := +by simp_rw [div_eq_mul_inv, h.mul_const] + +end normed_division_ring + +section is_R_or_C +variables {𝕜 : Type*} [is_R_or_C 𝕜] {f : α → 𝕜} + lemma integrable.of_real {f : α → ℝ} (hf : integrable f μ) : integrable (λ x, (f x : 𝕜)) μ := by { rw ← mem_ℒp_one_iff_integrable at hf ⊢, exact hf.of_real } From b9e46fe101fc897fb2e7edaf0bf1f09ea49eb81a Mon Sep 17 00:00:00 2001 From: Felix-Weilacher Date: Sat, 20 May 2023 16:28:32 +0000 Subject: [PATCH 1045/1291] refactor(topology): fix definition of residual (#18962) The current definition of `residual` in mathlib is incorrect for non-Baire spaces. This fixes it. Co-authored-by: Felix-Weilacher <112423742+Felix-Weilacher@users.noreply.github.com> --- src/order/filter/countable_Inter.lean | 71 ++++++++++++++++++++++- src/topology/G_delta.lean | 42 ++++++++++---- src/topology/metric_space/baire.lean | 83 ++++++++++----------------- 3 files changed, 132 insertions(+), 64 deletions(-) diff --git a/src/order/filter/countable_Inter.lean b/src/order/filter/countable_Inter.lean index 7c1b523d91fa1..9dbcb17230763 100644 --- a/src/order/filter/countable_Inter.lean +++ b/src/order/filter/countable_Inter.lean @@ -15,7 +15,7 @@ import data.set.countable In this file we define `countable_Inter_filter` to be the class of filters with the following property: for any countable collection of sets `s ∈ l` their intersection belongs to `l` as well. -Two main examples are the `residual` filter defined in `topology.metric_space.baire` and +Two main examples are the `residual` filter defined in `topology.G_delta` and the `measure.ae` filter defined in `measure_theory.measure_space`. We reformulate the definition in terms of indexed intersection and in terms of `filter.eventually` @@ -189,3 +189,72 @@ begin refine ⟨λ S hSc hS, ⟨_, _⟩⟩; refine (countable_sInter_mem hSc).2 (λ s hs, _), exacts [(hS s hs).1, (hS s hs).2] end + +namespace filter + +variable (g : set (set α)) + +/-- `filter.countable_generate_sets g` is the (sets of the) +greatest `countable_Inter_filter` containing `g`.-/ +inductive countable_generate_sets : set α → Prop +| basic {s : set α} : s ∈ g → countable_generate_sets s +| univ : countable_generate_sets univ +| superset {s t : set α} : countable_generate_sets s → s ⊆ t → countable_generate_sets t +| Inter {S : set (set α)} : S.countable → + (∀ s ∈ S, countable_generate_sets s) → countable_generate_sets ⋂₀ S + +/-- `filter.countable_generate g` is the greatest `countable_Inter_filter` containing `g`.-/ +@[derive countable_Inter_filter] +def countable_generate : filter α := +of_countable_Inter (countable_generate_sets g) (λ S, countable_generate_sets.Inter) + (λ s t, countable_generate_sets.superset) + +variable {g} + +/-- A set is in the `countable_Inter_filter` generated by `g` if and only if +it contains a countable intersection of elements of `g`. -/ +lemma mem_countable_generate_iff {s : set α} : s ∈ countable_generate g ↔ + ∃ (S : set (set α)), S ⊆ g ∧ S.countable ∧ ⋂₀ S ⊆ s := +begin + split; intro h, + { induction h with s hs s t hs st ih S Sct hS ih, + { exact ⟨{s}, by simp[hs]⟩ }, + { exact ⟨∅, by simp⟩ }, + { refine exists_imp_exists (λ S, _) ih, + tauto }, + choose T Tg Tct hT using ih, + refine ⟨⋃ s (H : s ∈ S), T s H, by simpa, Sct.bUnion Tct, _⟩, + apply subset_sInter, + intros s H, + refine subset_trans (sInter_subset_sInter (subset_Union₂ s H)) (hT s H), }, + rcases h with ⟨S, Sg, Sct, hS⟩, + refine mem_of_superset ((countable_sInter_mem Sct).mpr _) hS, + intros s H, + exact countable_generate_sets.basic (Sg H), +end + +lemma le_countable_generate_iff_of_countable_Inter_filter {f : filter α} + [countable_Inter_filter f] : f ≤ countable_generate g ↔ g ⊆ f.sets := +begin + split; intro h, + { exact subset_trans (λ s, countable_generate_sets.basic) h }, + intros s hs, + induction hs with s hs s t hs st ih S Sct hS ih, + { exact h hs }, + { exact univ_mem }, + { exact mem_of_superset ih st, }, + exact (countable_sInter_mem Sct).mpr ih, +end + +variable (g) + +/-- `countable_generate g` is the greatest `countable_Inter_filter` containing `g`.-/ +lemma countable_generate_is_greatest : is_greatest + {f : filter α | countable_Inter_filter f ∧ g ⊆ f.sets} (countable_generate g) := +begin + refine ⟨⟨infer_instance, λ s, countable_generate_sets.basic⟩, _⟩, + rintros f ⟨fct, hf⟩, + rwa @le_countable_generate_iff_of_countable_Inter_filter _ _ _ fct, +end + +end filter diff --git a/src/topology/G_delta.lean b/src/topology/G_delta.lean index d083170bcbafd..e67fa7e8bee47 100644 --- a/src/topology/G_delta.lean +++ b/src/topology/G_delta.lean @@ -5,6 +5,7 @@ Authors: Sébastien Gouëzel, Yury Kudryashov -/ import topology.uniform_space.basic import topology.separation +import order.filter.countable_Inter /-! # `Gδ` sets @@ -19,11 +20,8 @@ In this file we define `Gδ` sets and prove their basic properties. * `is_Gδ`: a set `s` is a `Gδ` set if it can be represented as an intersection of countably many open sets; -* `residual`: the filter of residual sets. A set `s` is called *residual* if it includes a dense - `Gδ` set. In a Baire space (e.g., in a complete (e)metric space), residual sets form a filter. - - For technical reasons, we define `residual` in any topological space but the definition agrees - with the description above only in Baire spaces. +* `residual`: the σ-filter of residual sets. A set `s` is called *residual* if it includes a + countable intersection of dense open sets. ## Main results @@ -180,10 +178,34 @@ end end continuous_at -/-- A set `s` is called *residual* if it includes a dense `Gδ` set. If `α` is a Baire space -(e.g., a complete metric space), then residual sets form a filter, see `mem_residual`. +section residual + +variable [topological_space α] -For technical reasons we define the filter `residual` in any topological space but in a non-Baire -space it is not useful because it may contain some non-residual sets. -/ +/-- A set `s` is called *residual* if it includes a countable intersection of dense open sets. -/ +@[derive countable_Inter_filter] def residual (α : Type*) [topological_space α] : filter α := -⨅ t (ht : is_Gδ t) (ht' : dense t), 𝓟 t +filter.countable_generate {t | is_open t ∧ dense t} + +instance countable_Inter_filter_residual : countable_Inter_filter (residual α) := +by rw [residual]; apply_instance + +/-- Dense open sets are residual. -/ +lemma residual_of_dense_open {s : set α} (ho : is_open s) (hd : dense s) : s ∈ residual α := +countable_generate_sets.basic ⟨ho, hd⟩ + +/-- Dense Gδ sets are residual. -/ +lemma residual_of_dense_Gδ {s : set α} (ho : is_Gδ s) (hd : dense s) : s ∈ residual α := +begin + rcases ho with ⟨T, To, Tct, rfl⟩, + exact (countable_sInter_mem Tct).mpr (λ t tT, residual_of_dense_open (To t tT) + (hd.mono (sInter_subset_of_mem tT))), +end + +/-- A set is residual iff it includes a countable intersection of dense open sets. -/ +lemma mem_residual_iff {s : set α} : s ∈ residual α ↔ + ∃ (S : set (set α)), (∀ t ∈ S, is_open t) ∧ (∀ t ∈ S, dense t) ∧ S.countable ∧ ⋂₀ S ⊆ s := +mem_countable_generate_iff.trans $ by simp_rw + [subset_def, mem_set_of, forall_and_distrib, and_assoc] + +end residual diff --git a/src/topology/metric_space/baire.lean b/src/topology/metric_space/baire.lean index 348298fedc1b2..21e4a43a3f940 100644 --- a/src/topology/metric_space/baire.lean +++ b/src/topology/metric_space/baire.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ import analysis.specific_limits.basic -import order.filter.countable_Inter import topology.G_delta import topology.sets.compacts @@ -22,8 +21,7 @@ intersection of dense Gδ sets is a dense Gδ set. We prove Baire theorem, givin formulations that can be handy. We also prove the important consequence that, if the space is covered by a countable union of closed sets, then the union of their interiors is dense. -We also define the filter `residual α` generated by dense `Gδ` sets and prove that this filter -has the countable intersection property. +We also prove that in Baire spaces, the `residual` sets are exactly those containing a dense Gδ set. -/ noncomputable theory @@ -222,25 +220,37 @@ begin { rwa forall_range_iff } end +/-- A set is residual (comeagre) if and only if it includes a dense `Gδ` set. -/ +lemma mem_residual {s : set α} : + s ∈ residual α ↔ ∃ t ⊆ s, is_Gδ t ∧ dense t := +begin + split, + { rw mem_residual_iff, + rintros ⟨S, hSo, hSd, Sct, Ss⟩, + refine ⟨_, Ss, ⟨_, λ t ht, hSo _ ht, Sct, rfl⟩, _⟩, + exact dense_sInter_of_open hSo Sct hSd, }, + rintros ⟨t, ts, ho, hd⟩, + exact mem_of_superset (residual_of_dense_Gδ ho hd) ts, +end + +/-- A property holds on a residual (comeagre) set if and only if it holds on some dense `Gδ` set. -/ +lemma eventually_residual {p : α → Prop} : + (∀ᶠ x in residual α, p x) ↔ ∃ (t : set α), is_Gδ t ∧ dense t ∧ ∀ (x : α), x ∈ t → p x := +begin + -- this can probably be improved... + convert @mem_residual _ _ _ p, + simp_rw [exists_prop, and_comm ((_ : set α) ⊆ p), and_assoc], + refl, +end + +lemma dense_of_mem_residual {s : set α} (hs : s ∈ residual α) : dense s := +let ⟨t, hts, _, hd⟩ := mem_residual.1 hs in hd.mono hts + /-- Baire theorem: a countable intersection of dense Gδ sets is dense. Formulated here with ⋂₀. -/ theorem dense_sInter_of_Gδ {S : set (set α)} (ho : ∀s∈S, is_Gδ s) (hS : S.countable) (hd : ∀s∈S, dense s) : dense (⋂₀S) := -begin - -- the result follows from the result for a countable intersection of dense open sets, - -- by rewriting each set as a countable intersection of open sets, which are of course dense. - choose T hTo hTc hsT using ho, - have : ⋂₀ S = ⋂₀ (⋃ s ∈ S, T s ‹_›), -- := (sInter_bUnion (λs hs, (hT s hs).2.2)).symm, - by simp only [sInter_Union, (hsT _ _).symm, ← sInter_eq_bInter], - rw this, - refine dense_sInter_of_open _ (hS.bUnion hTc) _; - simp only [mem_Union]; rintro t ⟨s, hs, tTs⟩, - show is_open t, from hTo s hs t tTs, - show dense t, - { intro x, - have := hd s hs x, - rw hsT s hs at this, - exact closure_mono (sInter_subset_of_mem tTs) this } -end +dense_of_mem_residual ((countable_sInter_mem hS).mpr + (λ s hs, residual_of_dense_Gδ (ho _ hs) (hd _ hs))) /-- Baire theorem: a countable intersection of dense Gδ sets is dense. Formulated here with an index set which is an encodable type. -/ @@ -270,39 +280,6 @@ begin apply dense_Inter_of_Gδ; simp [bool.forall_bool, *] end -/-- A property holds on a residual (comeagre) set if and only if it holds on some dense `Gδ` set. -/ -lemma eventually_residual {p : α → Prop} : - (∀ᶠ x in residual α, p x) ↔ ∃ (t : set α), is_Gδ t ∧ dense t ∧ ∀ x ∈ t, p x := -calc (∀ᶠ x in residual α, p x) ↔ - ∀ᶠ x in ⨅ (t : set α) (ht : is_Gδ t ∧ dense t), 𝓟 t, p x : - by simp only [residual, infi_and] -... ↔ ∃ (t : set α) (ht : is_Gδ t ∧ dense t), ∀ᶠ x in 𝓟 t, p x : mem_binfi_of_directed - (λ t₁ h₁ t₂ h₂, ⟨t₁ ∩ t₂, ⟨h₁.1.inter h₂.1, dense.inter_of_Gδ h₁.1 h₂.1 h₁.2 h₂.2⟩, by simp⟩) - ⟨univ, is_Gδ_univ, dense_univ⟩ -... ↔ _ : by simp [and_assoc] - -/-- A set is residual (comeagre) if and only if it includes a dense `Gδ` set. -/ -lemma mem_residual {s : set α} : - s ∈ residual α ↔ ∃ t ⊆ s, is_Gδ t ∧ dense t := -(@eventually_residual α _ _ (λ x, x ∈ s)).trans $ exists_congr $ -λ t, by rw [exists_prop, and_comm (t ⊆ s), subset_def, and_assoc] - -lemma dense_of_mem_residual {s : set α} (hs : s ∈ residual α) : - dense s := -let ⟨t, hts, _, hd⟩ := mem_residual.1 hs in hd.mono hts - -instance : countable_Inter_filter (residual α) := -⟨begin - intros S hSc hS, - simp only [mem_residual] at *, - choose T hTs hT using hS, - refine ⟨⋂ s ∈ S, T s ‹_›, _, _, _⟩, - { rw [sInter_eq_bInter], - exact Inter₂_mono hTs }, - { exact is_Gδ_bInter hSc (λ s hs, (hT s hs).1) }, - { exact dense_bInter_of_Gδ (λ s hs, (hT s hs).1) hSc (λ s hs, (hT s hs).2) } -end⟩ - /-- If a countable family of closed sets cover a dense `Gδ` set, then the union of their interiors is dense. Formulated here with `⋃`. -/ lemma is_Gδ.dense_Union_interior_of_closed [encodable ι] {s : set α} (hs : is_Gδ s) @@ -315,7 +292,7 @@ begin { refine dense_Inter_of_open hgo (λ i x, _), rw [closure_compl, interior_frontier (hc _)], exact id }, - refine (hd.inter_of_Gδ hs (is_Gδ_Inter $ λ i, (hgo i).is_Gδ) hgd).mono _, + refine (hd.inter_of_Gδ hs (is_Gδ_Inter_of_open $ λ i, hgo i) hgd).mono _, rintro x ⟨hxs, hxg⟩, rw [mem_Inter] at hxg, rcases mem_Union.1 (hU hxs) with ⟨i, hi⟩, From 33c67ae661dd8988516ff7f247b0be3018cdd952 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Sat, 20 May 2023 19:07:20 +0000 Subject: [PATCH 1046/1291] chore(*): add mathlib4 synchronization comments (#19042) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.category.Mon.limits` * `algebra.char_p.local_ring` * `algebra.direct_sum.decomposition` * `algebra.direct_sum.internal` * `algebra.module.bimodule` * `algebraic_topology.fundamental_groupoid.fundamental_group` * `algebraic_topology.fundamental_groupoid.punit` * `analysis.complex.operator_norm` * `analysis.normed_space.affine_isometry` * `analysis.normed_space.banach_steinhaus` * `analysis.normed_space.completion` * `analysis.normed_space.extend` * `analysis.normed_space.multilinear` * `analysis.special_functions.complex.arg` * `analysis.special_functions.complex.log` * `analysis.special_functions.pow.complex` * `category_theory.bicategory.free` * `category_theory.groupoid.free_groupoid` * `category_theory.monoidal.free.coherence` * `category_theory.monoidal.of_chosen_finite_products.basic` * `category_theory.monoidal.types.basic` * `data.matrix.kronecker` * `data.nat.factorial.double_factorial` * `data.polynomial.expand` * `linear_algebra.matrix.is_diag` * `linear_algebra.tensor_product.matrix` * `linear_algebra.unitary_group` * `number_theory.legendre_symbol.mul_character` * `number_theory.multiplicity` * `ring_theory.finite_presentation` * `ring_theory.ideal.local_ring` * `ring_theory.jacobson_ideal` * `ring_theory.localization.at_prime` * `ring_theory.matrix_algebra` * `ring_theory.nakayama` * `set_theory.game.pgame` * `topology.category.TopCommRing` * `topology.fiber_bundle.constructions` * `topology.sheaves.limits` * `topology.sheaves.presheaf` * `topology.sheaves.presheaf_of_functions` * `topology.sheaves.sheaf` --- src/algebra/category/Mon/limits.lean | 3 +++ src/algebra/char_p/local_ring.lean | 3 +++ src/algebra/direct_sum/decomposition.lean | 3 +++ src/algebra/direct_sum/internal.lean | 3 +++ src/algebra/module/bimodule.lean | 3 +++ .../fundamental_groupoid/fundamental_group.lean | 3 +++ src/algebraic_topology/fundamental_groupoid/punit.lean | 3 +++ src/analysis/complex/operator_norm.lean | 3 +++ src/analysis/normed_space/affine_isometry.lean | 3 +++ src/analysis/normed_space/banach_steinhaus.lean | 3 +++ src/analysis/normed_space/completion.lean | 3 +++ src/analysis/normed_space/extend.lean | 3 +++ src/analysis/normed_space/multilinear.lean | 3 +++ src/analysis/special_functions/complex/arg.lean | 3 +++ src/analysis/special_functions/complex/log.lean | 3 +++ src/analysis/special_functions/pow/complex.lean | 3 +++ src/category_theory/bicategory/free.lean | 3 +++ src/category_theory/groupoid/free_groupoid.lean | 3 +++ src/category_theory/monoidal/free/coherence.lean | 3 +++ .../monoidal/of_chosen_finite_products/basic.lean | 3 +++ src/category_theory/monoidal/types/basic.lean | 3 +++ src/data/matrix/kronecker.lean | 3 +++ src/data/nat/factorial/double_factorial.lean | 3 +++ src/data/polynomial/expand.lean | 3 +++ src/linear_algebra/matrix/is_diag.lean | 3 +++ src/linear_algebra/tensor_product/matrix.lean | 3 +++ src/linear_algebra/unitary_group.lean | 3 +++ src/number_theory/legendre_symbol/mul_character.lean | 3 +++ src/number_theory/multiplicity.lean | 3 +++ src/ring_theory/finite_presentation.lean | 3 +++ src/ring_theory/ideal/local_ring.lean | 3 +++ src/ring_theory/jacobson_ideal.lean | 3 +++ src/ring_theory/localization/at_prime.lean | 3 +++ src/ring_theory/matrix_algebra.lean | 3 +++ src/ring_theory/nakayama.lean | 3 +++ src/set_theory/game/pgame.lean | 3 +++ src/topology/category/TopCommRing.lean | 3 +++ src/topology/fiber_bundle/constructions.lean | 3 +++ src/topology/sheaves/limits.lean | 3 +++ src/topology/sheaves/presheaf.lean | 3 +++ src/topology/sheaves/presheaf_of_functions.lean | 3 +++ src/topology/sheaves/sheaf.lean | 3 +++ 42 files changed, 126 insertions(+) diff --git a/src/algebra/category/Mon/limits.lean b/src/algebra/category/Mon/limits.lean index 65297a1d52f7c..72d374a402eaa 100644 --- a/src/algebra/category/Mon/limits.lean +++ b/src/algebra/category/Mon/limits.lean @@ -12,6 +12,9 @@ import group_theory.submonoid.operations /-! # The category of (commutative) (additive) monoids has all limits +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Further, these limits are preserved by the forgetful functor --- that is, the underlying types are just the limits in the category of types. diff --git a/src/algebra/char_p/local_ring.lean b/src/algebra/char_p/local_ring.lean index d189ae2f98cf9..1d284e6a341d9 100644 --- a/src/algebra/char_p/local_ring.lean +++ b/src/algebra/char_p/local_ring.lean @@ -11,6 +11,9 @@ import data.nat.factorization.basic /-! # Characteristics of local rings +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main result - `char_p_zero_or_prime_power`: In a commutative local ring the characteristics is either diff --git a/src/algebra/direct_sum/decomposition.lean b/src/algebra/direct_sum/decomposition.lean index 86a1fe2c66343..de748982c5c8c 100644 --- a/src/algebra/direct_sum/decomposition.lean +++ b/src/algebra/direct_sum/decomposition.lean @@ -9,6 +9,9 @@ import algebra.module.submodule.basic /-! # Decompositions of additive monoids, groups, and modules into direct sums +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions * `direct_sum.decomposition ℳ`: A typeclass to provide a constructive decomposition from diff --git a/src/algebra/direct_sum/internal.lean b/src/algebra/direct_sum/internal.lean index 2ceb49d35e645..dc481b2926173 100644 --- a/src/algebra/direct_sum/internal.lean +++ b/src/algebra/direct_sum/internal.lean @@ -11,6 +11,9 @@ import algebra.direct_sum.algebra /-! # Internally graded rings and algebras +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This module provides `gsemiring` and `gcomm_semiring` instances for a collection of subobjects `A` when a `set_like.graded_monoid` instance is available: diff --git a/src/algebra/module/bimodule.lean b/src/algebra/module/bimodule.lean index 7c5acc489e8c9..259311e22dd0c 100644 --- a/src/algebra/module/bimodule.lean +++ b/src/algebra/module/bimodule.lean @@ -8,6 +8,9 @@ import ring_theory.tensor_product /-! # Bimodules +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + One frequently encounters situations in which several sets of scalars act on a single space, subject to compatibility condition(s). A distinguished instance of this is the theory of bimodules: one has two rings `R`, `S` acting on an additive group `M`, with `R` acting covariantly ("on the left") diff --git a/src/algebraic_topology/fundamental_groupoid/fundamental_group.lean b/src/algebraic_topology/fundamental_groupoid/fundamental_group.lean index 0589cbaeae128..1d8ccaa70cb48 100644 --- a/src/algebraic_topology/fundamental_groupoid/fundamental_group.lean +++ b/src/algebraic_topology/fundamental_groupoid/fundamental_group.lean @@ -12,6 +12,9 @@ import algebraic_topology.fundamental_groupoid.basic /-! # Fundamental group of a space +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given a topological space `X` and a basepoint `x`, the fundamental group is the automorphism group of `x` i.e. the group with elements being loops based at `x` (quotiented by homotopy equivalence). -/ diff --git a/src/algebraic_topology/fundamental_groupoid/punit.lean b/src/algebraic_topology/fundamental_groupoid/punit.lean index 118925d3fa84c..83bf875697b51 100644 --- a/src/algebraic_topology/fundamental_groupoid/punit.lean +++ b/src/algebraic_topology/fundamental_groupoid/punit.lean @@ -9,6 +9,9 @@ import algebraic_topology.fundamental_groupoid.basic /-! # Fundamental groupoid of punit +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The fundamental groupoid of punit is naturally isomorphic to `category_theory.discrete punit` -/ diff --git a/src/analysis/complex/operator_norm.lean b/src/analysis/complex/operator_norm.lean index a0133388f963b..c6637634392fd 100644 --- a/src/analysis/complex/operator_norm.lean +++ b/src/analysis/complex/operator_norm.lean @@ -9,6 +9,9 @@ import data.complex.determinant /-! # The basic continuous linear maps associated to `ℂ` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The continuous linear maps `complex.re_clm` (real part), `complex.im_clm` (imaginary part), `complex.conj_cle` (conjugation), and `complex.of_real_clm` (inclusion of `ℝ`) were introduced in `analysis.complex.operator_norm`. This file contains a few calculations requiring more imports: diff --git a/src/analysis/normed_space/affine_isometry.lean b/src/analysis/normed_space/affine_isometry.lean index c3e5a9e350546..5e6869ea196f4 100644 --- a/src/analysis/normed_space/affine_isometry.lean +++ b/src/analysis/normed_space/affine_isometry.lean @@ -12,6 +12,9 @@ import algebra.char_p.invertible /-! # Affine isometries +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define `affine_isometry 𝕜 P P₂` to be an affine isometric embedding of normed add-torsors `P` into `P₂` over normed `𝕜`-spaces and `affine_isometry_equiv` to be an affine isometric equivalence between `P` and `P₂`. diff --git a/src/analysis/normed_space/banach_steinhaus.lean b/src/analysis/normed_space/banach_steinhaus.lean index 215a4b76028b6..c8d8c11946514 100644 --- a/src/analysis/normed_space/banach_steinhaus.lean +++ b/src/analysis/normed_space/banach_steinhaus.lean @@ -9,6 +9,9 @@ import topology.algebra.module.basic /-! # The Banach-Steinhaus theorem: Uniform Boundedness Principle +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Herein we prove the Banach-Steinhaus theorem: any collection of bounded linear maps from a Banach space into a normed space which is pointwise bounded is uniformly bounded. diff --git a/src/analysis/normed_space/completion.lean b/src/analysis/normed_space/completion.lean index bac77e02d1122..d575ce664cc02 100644 --- a/src/analysis/normed_space/completion.lean +++ b/src/analysis/normed_space/completion.lean @@ -10,6 +10,9 @@ import topology.algebra.uniform_ring /-! # Normed space structure on the completion of a normed space +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + If `E` is a normed space over `𝕜`, then so is `uniform_space.completion E`. In this file we provide necessary instances and define `uniform_space.completion.to_complₗᵢ` - coercion `E → uniform_space.completion E` as a bundled linear isometry. diff --git a/src/analysis/normed_space/extend.lean b/src/analysis/normed_space/extend.lean index a9080f14c606e..1b12d11cc1e0c 100644 --- a/src/analysis/normed_space/extend.lean +++ b/src/analysis/normed_space/extend.lean @@ -11,6 +11,9 @@ import data.is_R_or_C.basic /-! # Extending a continuous `ℝ`-linear map to a continuous `𝕜`-linear map +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we provide a way to extend a continuous `ℝ`-linear map to a continuous `𝕜`-linear map in a way that bounds the norm by the norm of the original map, when `𝕜` is either `ℝ` (the extension is trivial) or `ℂ`. We formulate the extension uniformly, by assuming `is_R_or_C 𝕜`. diff --git a/src/analysis/normed_space/multilinear.lean b/src/analysis/normed_space/multilinear.lean index 73546f022c1f3..da305f925b08a 100644 --- a/src/analysis/normed_space/multilinear.lean +++ b/src/analysis/normed_space/multilinear.lean @@ -9,6 +9,9 @@ import topology.algebra.module.multilinear /-! # Operator norm on the space of continuous multilinear maps +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + When `f` is a continuous multilinear map in finitely many variables, we define its norm `‖f‖` as the smallest number such that `‖f m‖ ≤ ‖f‖ * ∏ i, ‖m i‖` for all `m`. diff --git a/src/analysis/special_functions/complex/arg.lean b/src/analysis/special_functions/complex/arg.lean index 28ee90e5089ed..e9c916006e5c4 100644 --- a/src/analysis/special_functions/complex/arg.lean +++ b/src/analysis/special_functions/complex/arg.lean @@ -9,6 +9,9 @@ import analysis.special_functions.trigonometric.inverse /-! # The argument of a complex number. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define `arg : ℂ → ℝ`, returing a real number in the range (-π, π], such that for `x ≠ 0`, `sin (arg x) = x.im / x.abs` and `cos (arg x) = x.re / x.abs`, while `arg 0` defaults to `0` diff --git a/src/analysis/special_functions/complex/log.lean b/src/analysis/special_functions/complex/log.lean index 5a55d5b154cfc..7965cd052fc0a 100644 --- a/src/analysis/special_functions/complex/log.lean +++ b/src/analysis/special_functions/complex/log.lean @@ -9,6 +9,9 @@ import analysis.special_functions.log.basic /-! # The complex `log` function +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Basic properties, relationship with `exp`. -/ diff --git a/src/analysis/special_functions/pow/complex.lean b/src/analysis/special_functions/pow/complex.lean index d8a42ea6f78a9..9a4fc47cb086b 100644 --- a/src/analysis/special_functions/pow/complex.lean +++ b/src/analysis/special_functions/pow/complex.lean @@ -8,6 +8,9 @@ import analysis.special_functions.complex.log /-! # Power function on `ℂ` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We construct the power functions `x ^ y`, where `x` and `y` are complex numbers. -/ diff --git a/src/category_theory/bicategory/free.lean b/src/category_theory/bicategory/free.lean index 23ace75977b00..a47147dbff248 100644 --- a/src/category_theory/bicategory/free.lean +++ b/src/category_theory/bicategory/free.lean @@ -8,6 +8,9 @@ import category_theory.bicategory.functor /-! # Free bicategories +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define the free bicategory over a quiver. In this bicategory, the 1-morphisms are freely generated by the arrows in the quiver, and the 2-morphisms are freely generated by the formal identities, the formal unitors, and the formal associators modulo the relation derived from the diff --git a/src/category_theory/groupoid/free_groupoid.lean b/src/category_theory/groupoid/free_groupoid.lean index 4a683ef2d8107..33fb79fc5f491 100644 --- a/src/category_theory/groupoid/free_groupoid.lean +++ b/src/category_theory/groupoid/free_groupoid.lean @@ -14,6 +14,9 @@ import combinatorics.quiver.symmetric /-! # Free groupoid on a quiver +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the free groupoid on a quiver, the lifting of a prefunctor to its unique extension as a functor from the free groupoid, and proves uniqueness of this extension. diff --git a/src/category_theory/monoidal/free/coherence.lean b/src/category_theory/monoidal/free/coherence.lean index f2fdac8bbe094..8aab1c0a2e6f7 100644 --- a/src/category_theory/monoidal/free/coherence.lean +++ b/src/category_theory/monoidal/free/coherence.lean @@ -10,6 +10,9 @@ import category_theory.discrete_category /-! # The monoidal coherence theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we prove the monoidal coherence theorem, stated in the following form: the free monoidal category over any type `C` is thin. diff --git a/src/category_theory/monoidal/of_chosen_finite_products/basic.lean b/src/category_theory/monoidal/of_chosen_finite_products/basic.lean index 93e166e639cc6..fb1123060e82a 100644 --- a/src/category_theory/monoidal/of_chosen_finite_products/basic.lean +++ b/src/category_theory/monoidal/of_chosen_finite_products/basic.lean @@ -10,6 +10,9 @@ import category_theory.pempty /-! # The monoidal structure on a category with chosen finite products. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This is a variant of the development in `category_theory.monoidal.of_has_finite_products`, which uses specified choices of the terminal object and binary product, enabling the construction of a cartesian category with specific definitions of the tensor unit diff --git a/src/category_theory/monoidal/types/basic.lean b/src/category_theory/monoidal/types/basic.lean index 84bcfd68e41f7..72904c52b2636 100644 --- a/src/category_theory/monoidal/types/basic.lean +++ b/src/category_theory/monoidal/types/basic.lean @@ -10,6 +10,9 @@ import logic.equiv.fin /-! # The category of types is a monoidal category + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ open category_theory diff --git a/src/data/matrix/kronecker.lean b/src/data/matrix/kronecker.lean index 9c07d87b561bf..761a437a69959 100644 --- a/src/data/matrix/kronecker.lean +++ b/src/data/matrix/kronecker.lean @@ -14,6 +14,9 @@ import ring_theory.tensor_product /-! # Kronecker product of matrices +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This defines the [Kronecker product](https://en.wikipedia.org/wiki/Kronecker_product). ## Main definitions diff --git a/src/data/nat/factorial/double_factorial.lean b/src/data/nat/factorial/double_factorial.lean index 87b621056a376..5964c0ef8dd6a 100644 --- a/src/data/nat/factorial/double_factorial.lean +++ b/src/data/nat/factorial/double_factorial.lean @@ -9,6 +9,9 @@ import tactic.ring /-! # Double factorials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the double factorial, `n‼ := n * (n - 2) * (n - 4) * ...`. diff --git a/src/data/polynomial/expand.lean b/src/data/polynomial/expand.lean index 0ea5919000489..8df70d7d9590a 100644 --- a/src/data/polynomial/expand.lean +++ b/src/data/polynomial/expand.lean @@ -10,6 +10,9 @@ import tactic.ring_exp /-! # Expand a polynomial by a factor of p, so `∑ aₙ xⁿ` becomes `∑ aₙ xⁿᵖ`. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions * `polynomial.expand R p f`: expand the polynomial `f` with coefficients in a diff --git a/src/linear_algebra/matrix/is_diag.lean b/src/linear_algebra/matrix/is_diag.lean index 11c1ab8f5962c..4616127392c31 100644 --- a/src/linear_algebra/matrix/is_diag.lean +++ b/src/linear_algebra/matrix/is_diag.lean @@ -10,6 +10,9 @@ import data.matrix.kronecker /-! # Diagonal matrices +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains the definition and basic results about diagonal matrices. ## Main results diff --git a/src/linear_algebra/tensor_product/matrix.lean b/src/linear_algebra/tensor_product/matrix.lean index a4490a52c0edc..9b97cdba60234 100644 --- a/src/linear_algebra/tensor_product/matrix.lean +++ b/src/linear_algebra/tensor_product/matrix.lean @@ -11,6 +11,9 @@ import linear_algebra.tensor_product_basis /-! # Connections between `tensor_product` and `matrix` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains results about the matrices corresponding to maps between tensor product types, where the correspondance is induced by `basis.tensor_product` diff --git a/src/linear_algebra/unitary_group.lean b/src/linear_algebra/unitary_group.lean index 9a72696fc3684..dac05f32d4cf5 100644 --- a/src/linear_algebra/unitary_group.lean +++ b/src/linear_algebra/unitary_group.lean @@ -11,6 +11,9 @@ import algebra.star.unitary /-! # The Unitary Group +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines elements of the unitary group `unitary_group n α`, where `α` is a `star_ring`. This consists of all `n` by `n` matrices with entries in `α` such that the star-transpose is its inverse. In addition, we define the group structure on `unitary_group n α`, and the embedding into diff --git a/src/number_theory/legendre_symbol/mul_character.lean b/src/number_theory/legendre_symbol/mul_character.lean index e3c827cc2fb1b..3fa2fd2b10673 100644 --- a/src/number_theory/legendre_symbol/mul_character.lean +++ b/src/number_theory/legendre_symbol/mul_character.lean @@ -10,6 +10,9 @@ import data.fintype.units /-! # Multiplicative characters of finite rings and fields +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Let `R` and `R'` be a commutative rings. A *multiplicative character* of `R` with values in `R'` is a morphism of monoids from the multiplicative monoid of `R` into that of `R'` diff --git a/src/number_theory/multiplicity.lean b/src/number_theory/multiplicity.lean index 5823e572d478f..8231f6de58a91 100644 --- a/src/number_theory/multiplicity.lean +++ b/src/number_theory/multiplicity.lean @@ -12,6 +12,9 @@ import ring_theory.ideal.quotient_operations /-! # Multiplicity in Number Theory +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains results in number theory relating to multiplicity. ## Main statements diff --git a/src/ring_theory/finite_presentation.lean b/src/ring_theory/finite_presentation.lean index 3084be7a24354..87e17d5d983e2 100644 --- a/src/ring_theory/finite_presentation.lean +++ b/src/ring_theory/finite_presentation.lean @@ -11,6 +11,9 @@ import ring_theory.ideal.quotient_operations /-! # Finiteness conditions in commutative algebra +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define several notions of finiteness that are common in commutative algebra. ## Main declarations diff --git a/src/ring_theory/ideal/local_ring.lean b/src/ring_theory/ideal/local_ring.lean index 887c846030036..ef24d74567017 100644 --- a/src/ring_theory/ideal/local_ring.lean +++ b/src/ring_theory/ideal/local_ring.lean @@ -13,6 +13,9 @@ import logic.equiv.transfer_instance # Local rings +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Define local rings as commutative rings having a unique maximal ideal. ## Main definitions diff --git a/src/ring_theory/jacobson_ideal.lean b/src/ring_theory/jacobson_ideal.lean index 0e1f0e5efbbcb..85d0476833875 100644 --- a/src/ring_theory/jacobson_ideal.lean +++ b/src/ring_theory/jacobson_ideal.lean @@ -9,6 +9,9 @@ import ring_theory.polynomial.quotient /-! # Jacobson radical +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The Jacobson radical of a ring `R` is defined to be the intersection of all maximal ideals of `R`. This is similar to how the nilradical is equal to the intersection of all prime ideals of `R`. diff --git a/src/ring_theory/localization/at_prime.lean b/src/ring_theory/localization/at_prime.lean index 1b3a5bc1a1418..fc7325ff90e02 100644 --- a/src/ring_theory/localization/at_prime.lean +++ b/src/ring_theory/localization/at_prime.lean @@ -9,6 +9,9 @@ import ring_theory.localization.ideal /-! # Localizations of commutative rings at the complement of a prime ideal +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions * `is_localization.at_prime (I : ideal R) [is_prime I] (S : Type*)` expresses that `S` is a diff --git a/src/ring_theory/matrix_algebra.lean b/src/ring_theory/matrix_algebra.lean index 323fc9aae3d16..11c41cba4fd22 100644 --- a/src/ring_theory/matrix_algebra.lean +++ b/src/ring_theory/matrix_algebra.lean @@ -7,6 +7,9 @@ import data.matrix.basis import ring_theory.tensor_product /-! +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We show `matrix n n A ≃ₐ[R] (A ⊗[R] matrix n n R)`. -/ diff --git a/src/ring_theory/nakayama.lean b/src/ring_theory/nakayama.lean index e1ee04629f7c5..b5fe642599314 100644 --- a/src/ring_theory/nakayama.lean +++ b/src/ring_theory/nakayama.lean @@ -8,6 +8,9 @@ import ring_theory.jacobson_ideal /-! # Nakayama's lemma +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains some alternative statements of Nakayama's Lemma as found in [Stacks: Nakayama's Lemma](https://stacks.math.columbia.edu/tag/00DV). diff --git a/src/set_theory/game/pgame.lean b/src/set_theory/game/pgame.lean index fd8e873590ce1..1cfa0e55c037c 100644 --- a/src/set_theory/game/pgame.lean +++ b/src/set_theory/game/pgame.lean @@ -11,6 +11,9 @@ import order.game_add /-! # Combinatorial (pre-)games. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The basic theory of combinatorial games, following Conway's book `On Numbers and Games`. We construct "pregames", define an ordering and arithmetic operations on them, then show that the operations descend to "games", defined via the equivalence relation `p ≈ q ↔ p ≤ q ∧ q ≤ p`. diff --git a/src/topology/category/TopCommRing.lean b/src/topology/category/TopCommRing.lean index 7a2813a07a1d9..a1545ad6343de 100644 --- a/src/topology/category/TopCommRing.lean +++ b/src/topology/category/TopCommRing.lean @@ -10,6 +10,9 @@ import topology.algebra.ring.basic /-! # Category of topological commutative rings +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We introduce the category `TopCommRing` of topological commutative rings together with the relevant forgetful functors to topological spaces and commutative rings. -/ diff --git a/src/topology/fiber_bundle/constructions.lean b/src/topology/fiber_bundle/constructions.lean index fe7aca58bf792..4607bf248af1c 100644 --- a/src/topology/fiber_bundle/constructions.lean +++ b/src/topology/fiber_bundle/constructions.lean @@ -8,6 +8,9 @@ import topology.fiber_bundle.basic /-! # Standard constructions on fiber bundles +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains several standard constructions on fiber bundles: * `bundle.trivial.fiber_bundle 𝕜 B F`: the trivial fiber bundle with model fiber `F` over the base diff --git a/src/topology/sheaves/limits.lean b/src/topology/sheaves/limits.lean index 90cbca28a2112..79f129cd7ce4b 100644 --- a/src/topology/sheaves/limits.lean +++ b/src/topology/sheaves/limits.lean @@ -9,6 +9,9 @@ import category_theory.limits.functor_category /-! # Presheaves in `C` have limits and colimits when `C` does. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ noncomputable theory diff --git a/src/topology/sheaves/presheaf.lean b/src/topology/sheaves/presheaf.lean index 6e63c82002e36..e4c312be03265 100644 --- a/src/topology/sheaves/presheaf.lean +++ b/src/topology/sheaves/presheaf.lean @@ -10,6 +10,9 @@ import category_theory.adjunction.opposites /-! # Presheaves on a topological space +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define `presheaf C X` simply as `(opens X)ᵒᵖ ⥤ C`, and inherit the category structure with natural transformations as morphisms. diff --git a/src/topology/sheaves/presheaf_of_functions.lean b/src/topology/sheaves/presheaf_of_functions.lean index 216bb9f08f765..07212f6643439 100644 --- a/src/topology/sheaves/presheaf_of_functions.lean +++ b/src/topology/sheaves/presheaf_of_functions.lean @@ -11,6 +11,9 @@ import topology.continuous_function.algebra /-! # Presheaves of functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We construct some simple examples of presheaves of functions on a topological space. * `presheaf_to_Types X T`, where `T : X → Type`, is the presheaf of dependently-typed (not-necessarily continuous) functions diff --git a/src/topology/sheaves/sheaf.lean b/src/topology/sheaves/sheaf.lean index 46ce853717d86..4c0ccff469017 100644 --- a/src/topology/sheaves/sheaf.lean +++ b/src/topology/sheaves/sheaf.lean @@ -10,6 +10,9 @@ import category_theory.sites.spaces /-! # Sheaves +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define sheaves on a topological space, with values in an arbitrary category. A presheaf on a topological space `X` is a sheaf presicely when it is a sheaf under the From 1b0a28e1c93409dbf6d69526863cd9984ef652ce Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Sun, 21 May 2023 06:27:02 +0000 Subject: [PATCH 1047/1291] chore(*): add mathlib4 synchronization comments (#19055) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.category.Module.projective` * `algebra.monoid_algebra.to_direct_sum` * `analysis.convex.strict_convex_space` * `analysis.convex.uniform` * `analysis.normed_space.banach` * `analysis.normed_space.bounded_linear_maps` * `analysis.normed_space.complemented` * `analysis.normed_space.finite_dimension` * `analysis.normed_space.mazur_ulam` * `analysis.special_functions.log.base` * `analysis.special_functions.log.monotone` * `analysis.special_functions.pow.real` * `combinatorics.additive.salem_spencer` * `data.is_R_or_C.lemmas` * `linear_algebra.adic_completion` * `linear_algebra.cross_product` * `measure_theory.measure.doubling` * `ring_theory.graded_algebra.basic` * `ring_theory.mv_polynomial.homogeneous` * `ring_theory.valuation.extend_to_localization` * `topology.algebra.module.star` --- src/algebra/category/Module/projective.lean | 3 +++ src/algebra/monoid_algebra/to_direct_sum.lean | 3 +++ src/analysis/convex/strict_convex_space.lean | 3 +++ src/analysis/convex/uniform.lean | 3 +++ src/analysis/normed_space/banach.lean | 3 +++ src/analysis/normed_space/bounded_linear_maps.lean | 3 +++ src/analysis/normed_space/complemented.lean | 3 +++ src/analysis/normed_space/finite_dimension.lean | 3 +++ src/analysis/normed_space/mazur_ulam.lean | 3 +++ src/analysis/special_functions/log/base.lean | 3 +++ src/analysis/special_functions/log/monotone.lean | 3 +++ src/analysis/special_functions/pow/real.lean | 3 +++ src/combinatorics/additive/salem_spencer.lean | 3 +++ src/data/is_R_or_C/lemmas.lean | 5 ++++- src/linear_algebra/adic_completion.lean | 3 +++ src/linear_algebra/cross_product.lean | 3 +++ src/measure_theory/measure/doubling.lean | 3 +++ src/ring_theory/graded_algebra/basic.lean | 3 +++ src/ring_theory/mv_polynomial/homogeneous.lean | 3 +++ src/ring_theory/valuation/extend_to_localization.lean | 3 +++ src/topology/algebra/module/star.lean | 3 +++ 21 files changed, 64 insertions(+), 1 deletion(-) diff --git a/src/algebra/category/Module/projective.lean b/src/algebra/category/Module/projective.lean index 4ae2bb3bf1898..6fa327c3db2be 100644 --- a/src/algebra/category/Module/projective.lean +++ b/src/algebra/category/Module/projective.lean @@ -10,6 +10,9 @@ import linear_algebra.finsupp_vector_space /-! # The category of `R`-modules has enough projectives. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ universes v u diff --git a/src/algebra/monoid_algebra/to_direct_sum.lean b/src/algebra/monoid_algebra/to_direct_sum.lean index 4f43e85416d37..4a42590b85905 100644 --- a/src/algebra/monoid_algebra/to_direct_sum.lean +++ b/src/algebra/monoid_algebra/to_direct_sum.lean @@ -10,6 +10,9 @@ import data.finsupp.to_dfinsupp /-! # Conversion between `add_monoid_algebra` and homogenous `direct_sum` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This module provides conversions between `add_monoid_algebra` and `direct_sum`. The latter is essentially a dependent version of the former. diff --git a/src/analysis/convex/strict_convex_space.lean b/src/analysis/convex/strict_convex_space.lean index 751345c147d10..189e75c6b3c98 100644 --- a/src/analysis/convex/strict_convex_space.lean +++ b/src/analysis/convex/strict_convex_space.lean @@ -13,6 +13,9 @@ import analysis.normed_space.affine_isometry /-! # Strictly convex spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines strictly convex spaces. A normed space is strictly convex if all closed balls are strictly convex. This does **not** mean that the norm is strictly convex (in fact, it never is). diff --git a/src/analysis/convex/uniform.lean b/src/analysis/convex/uniform.lean index 5a7fc0b019ddf..e8c1bab9550df 100644 --- a/src/analysis/convex/uniform.lean +++ b/src/analysis/convex/uniform.lean @@ -8,6 +8,9 @@ import analysis.convex.strict_convex_space /-! # Uniformly convex spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines uniformly convex spaces, which are real normed vector spaces in which for all strictly positive `ε`, there exists some strictly positive `δ` such that `ε ≤ ‖x - y‖` implies `‖x + y‖ ≤ 2 - δ` for all `x` and `y` of norm at most than `1`. This means that the triangle diff --git a/src/analysis/normed_space/banach.lean b/src/analysis/normed_space/banach.lean index 83d979420c9ad..e3aa13499386f 100644 --- a/src/analysis/normed_space/banach.lean +++ b/src/analysis/normed_space/banach.lean @@ -10,6 +10,9 @@ import analysis.normed_space.affine_isometry /-! # Banach open mapping theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains the Banach open mapping theorem, i.e., the fact that a bijective bounded linear map between Banach spaces has a bounded inverse. -/ diff --git a/src/analysis/normed_space/bounded_linear_maps.lean b/src/analysis/normed_space/bounded_linear_maps.lean index d7f734a30866d..5f3aeef6d4354 100644 --- a/src/analysis/normed_space/bounded_linear_maps.lean +++ b/src/analysis/normed_space/bounded_linear_maps.lean @@ -10,6 +10,9 @@ import analysis.asymptotics.asymptotics /-! # Bounded linear maps +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines a class stating that a map between normed vector spaces is (bi)linear and continuous. Instead of asking for continuity, the definition takes the equivalent condition (because the space diff --git a/src/analysis/normed_space/complemented.lean b/src/analysis/normed_space/complemented.lean index 0ab67946b023c..7c71577bec79b 100644 --- a/src/analysis/normed_space/complemented.lean +++ b/src/analysis/normed_space/complemented.lean @@ -9,6 +9,9 @@ import analysis.normed_space.finite_dimension /-! # Complemented subspaces of normed vector spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A submodule `p` of a topological module `E` over `R` is called *complemented* if there exists a continuous linear projection `f : E →ₗ[R] p`, `∀ x : p, f x = x`. We prove that for a closed subspace of a normed space this condition is equivalent to existence of a closed diff --git a/src/analysis/normed_space/finite_dimension.lean b/src/analysis/normed_space/finite_dimension.lean index b24cf93f7ea0b..9739b24088220 100644 --- a/src/analysis/normed_space/finite_dimension.lean +++ b/src/analysis/normed_space/finite_dimension.lean @@ -15,6 +15,9 @@ import topology.instances.matrix /-! # Finite dimensional normed spaces over complete fields +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Over a complete nontrivially normed field, in finite dimension, all norms are equivalent and all linear maps are continuous. Moreover, a finite-dimensional subspace is always complete and closed. diff --git a/src/analysis/normed_space/mazur_ulam.lean b/src/analysis/normed_space/mazur_ulam.lean index ff8c8ce673de2..4a0ac2d03bf05 100644 --- a/src/analysis/normed_space/mazur_ulam.lean +++ b/src/analysis/normed_space/mazur_ulam.lean @@ -9,6 +9,9 @@ import analysis.normed_space.affine_isometry /-! # Mazur-Ulam Theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Mazur-Ulam theorem states that an isometric bijection between two normed affine spaces over `ℝ` is affine. We formalize it in three definitions: diff --git a/src/analysis/special_functions/log/base.lean b/src/analysis/special_functions/log/base.lean index 03073d7be214a..02b23f546daa7 100644 --- a/src/analysis/special_functions/log/base.lean +++ b/src/analysis/special_functions/log/base.lean @@ -9,6 +9,9 @@ import data.int.log /-! # Real logarithm base `b` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define `real.logb` to be the logarithm of a real number in a given base `b`. We define this as the division of the natural logarithms of the argument and the base, so that we have a globally defined function with `logb b 0 = 0`, `logb b (-x) = logb b x` `logb 0 x = 0` and diff --git a/src/analysis/special_functions/log/monotone.lean b/src/analysis/special_functions/log/monotone.lean index 9eb14b1c14252..8e248092cccd1 100644 --- a/src/analysis/special_functions/log/monotone.lean +++ b/src/analysis/special_functions/log/monotone.lean @@ -8,6 +8,9 @@ import analysis.special_functions.pow.real /-! # Logarithm Tonality +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we describe the tonality of the logarithm function when multiplied by functions of the form `x ^ a`. diff --git a/src/analysis/special_functions/pow/real.lean b/src/analysis/special_functions/pow/real.lean index 6815e879fee35..26582880834d2 100644 --- a/src/analysis/special_functions/pow/real.lean +++ b/src/analysis/special_functions/pow/real.lean @@ -8,6 +8,9 @@ import analysis.special_functions.pow.complex /-! # Power function on `ℝ` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We construct the power functions `x ^ y`, where `x` and `y` are real numbers. -/ diff --git a/src/combinatorics/additive/salem_spencer.lean b/src/combinatorics/additive/salem_spencer.lean index e7b8bfc1f8836..7b6fc3d940e27 100644 --- a/src/combinatorics/additive/salem_spencer.lean +++ b/src/combinatorics/additive/salem_spencer.lean @@ -10,6 +10,9 @@ import analysis.convex.strict_convex_space /-! # Salem-Spencer sets and Roth numbers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines Salem-Spencer sets and the Roth number of a set. A Salem-Spencer set is a set without arithmetic progressions of length `3`. Equivalently, the diff --git a/src/data/is_R_or_C/lemmas.lean b/src/data/is_R_or_C/lemmas.lean index f4685686068ff..93d976cb3acab 100644 --- a/src/data/is_R_or_C/lemmas.lean +++ b/src/data/is_R_or_C/lemmas.lean @@ -7,7 +7,10 @@ import analysis.normed_space.finite_dimension import field_theory.tower import data.is_R_or_C.basic -/-! # Further lemmas about `is_R_or_C` -/ +/-! # Further lemmas about `is_R_or_C` + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4.-/ variables {K E : Type*} [is_R_or_C K] diff --git a/src/linear_algebra/adic_completion.lean b/src/linear_algebra/adic_completion.lean index bc42c89538b27..3f73138ad4edd 100644 --- a/src/linear_algebra/adic_completion.lean +++ b/src/linear_algebra/adic_completion.lean @@ -11,6 +11,9 @@ import ring_theory.jacobson_ideal /-! # Completion of a module with respect to an ideal. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define the notions of Hausdorff, precomplete, and complete for an `R`-module `M` with respect to an ideal `I`: diff --git a/src/linear_algebra/cross_product.lean b/src/linear_algebra/cross_product.lean index 67f47a6f8ec05..524459f749110 100644 --- a/src/linear_algebra/cross_product.lean +++ b/src/linear_algebra/cross_product.lean @@ -12,6 +12,9 @@ import algebra.lie.basic /-! # Cross products +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This module defines the cross product of vectors in $R^3$ for $R$ a commutative ring, as a bilinear map. diff --git a/src/measure_theory/measure/doubling.lean b/src/measure_theory/measure/doubling.lean index 1ca5de2930d7d..c306deac59ea1 100644 --- a/src/measure_theory/measure/doubling.lean +++ b/src/measure_theory/measure/doubling.lean @@ -9,6 +9,9 @@ import measure_theory.measure.measure_space_def /-! # Uniformly locally doubling measures +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A uniformly locally doubling measure `μ` on a metric space is a measure for which there exists a constant `C` such that for all sufficiently small radii `ε`, and for any centre, the measure of a ball of radius `2 * ε` is bounded by `C` times the measure of the concentric ball of radius `ε`. diff --git a/src/ring_theory/graded_algebra/basic.lean b/src/ring_theory/graded_algebra/basic.lean index 1e958d27d4ea1..86eebeb51333a 100644 --- a/src/ring_theory/graded_algebra/basic.lean +++ b/src/ring_theory/graded_algebra/basic.lean @@ -11,6 +11,9 @@ import algebra.direct_sum.ring /-! # Internally-graded rings and algebras +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the typeclass `graded_algebra 𝒜`, for working with an algebra `A` that is internally graded by a collection of submodules `𝒜 : ι → submodule R A`. See the docstring of that typeclass for more information. diff --git a/src/ring_theory/mv_polynomial/homogeneous.lean b/src/ring_theory/mv_polynomial/homogeneous.lean index 40f0e3987b7ce..e3b3e6e53cbd8 100644 --- a/src/ring_theory/mv_polynomial/homogeneous.lean +++ b/src/ring_theory/mv_polynomial/homogeneous.lean @@ -10,6 +10,9 @@ import data.mv_polynomial.variables /-! # Homogeneous polynomials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A multivariate polynomial `φ` is homogeneous of degree `n` if all monomials occuring in `φ` have degree `n`. diff --git a/src/ring_theory/valuation/extend_to_localization.lean b/src/ring_theory/valuation/extend_to_localization.lean index 03c7c8674849d..6175b6f334b17 100644 --- a/src/ring_theory/valuation/extend_to_localization.lean +++ b/src/ring_theory/valuation/extend_to_localization.lean @@ -10,6 +10,9 @@ import ring_theory.valuation.basic # Extending valuations to a localization +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We show that, given a valuation `v` taking values in a linearly ordered commutative *group* with zero `Γ`, and a submonoid `S` of `v.supp.prime_compl`, the valuation `v` can be naturally extended to the localization `S⁻¹A`. diff --git a/src/topology/algebra/module/star.lean b/src/topology/algebra/module/star.lean index 314d64d6b779c..07074a3d670d8 100644 --- a/src/topology/algebra/module/star.lean +++ b/src/topology/algebra/module/star.lean @@ -9,6 +9,9 @@ import topology.algebra.star /-! # The star operation, bundled as a continuous star-linear equiv + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ /-- If `A` is a topological module over a commutative `R` with compatible actions, From b2c89893177f66a48daf993b7ba5ef7cddeff8c9 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sun, 21 May 2023 20:10:14 +0000 Subject: [PATCH 1048/1291] chore(data/multiset/pi): correct names and reorder (#19050) * `multiset.pi_cons_injective` is about `multiset.pi.cons` so should have a `.` in its name. * `multiset.pi.cons_ext` is not an ext lemma, but more closely resembles eta-reduction. This also groups together the lemmas about `pi.cons`. --- src/algebra/big_operators/ring.lean | 2 +- src/data/finset/pi.lean | 6 ++--- src/data/multiset/pi.lean | 42 ++++++++++++++--------------- 3 files changed, 25 insertions(+), 25 deletions(-) diff --git a/src/algebra/big_operators/ring.lean b/src/algebra/big_operators/ring.lean index 8ed6121b1b0d7..6a0c5a9dca0dd 100644 --- a/src/algebra/big_operators/ring.lean +++ b/src/algebra/big_operators/ring.lean @@ -101,7 +101,7 @@ begin rw [prod_insert ha, pi_insert ha, ih, sum_mul, sum_bUnion h₁], refine sum_congr rfl (λ b _, _), have h₂ : ∀p₁∈pi s t, ∀p₂∈pi s t, pi.cons s a b p₁ = pi.cons s a b p₂ → p₁ = p₂, from - assume p₁ h₁ p₂ h₂ eq, pi_cons_injective ha eq, + assume p₁ h₁ p₂ h₂ eq, pi.cons_injective ha eq, rw [sum_image h₂, mul_sum], refine sum_congr rfl (λ g _, _), rw [attach_insert, prod_insert, prod_image], diff --git a/src/data/finset/pi.lean b/src/data/finset/pi.lean index f5ebe2ba95aa6..0978ffba703e9 100644 --- a/src/data/finset/pi.lean +++ b/src/data/finset/pi.lean @@ -58,10 +58,10 @@ lemma pi.cons_ne {s : finset α} {a a' : α} {b : δ a} {f : Πa, a ∈ s → δ pi.cons s a b f a' h = f a' ((mem_insert.1 h).resolve_left ha.symm) := multiset.pi.cons_ne _ _ -lemma pi_cons_injective {a : α} {b : δ a} {s : finset α} (hs : a ∉ s) : +lemma pi.cons_injective {a : α} {b : δ a} {s : finset α} (hs : a ∉ s) : function.injective (pi.cons s a b) := assume e₁ e₂ eq, -@multiset.pi_cons_injective α _ δ a b s.1 hs _ _ $ +@multiset.pi.cons_injective α _ δ a b s.1 hs _ _ $ funext $ assume e, funext $ assume h, have pi.cons s a b e₁ e (by simpa only [multiset.mem_cons, mem_insert] using h) = pi.cons s a b e₂ e (by simpa only [multiset.mem_cons, mem_insert] using h), @@ -84,7 +84,7 @@ begin subst s', rw pi_cons, congr, funext b, refine ((pi s t).nodup.map _).dedup.symm, - exact multiset.pi_cons_injective ha, + exact multiset.pi.cons_injective ha, end lemma pi_singletons {β : Type*} (s : finset α) (f : α → β) : diff --git a/src/data/multiset/pi.lean b/src/data/multiset/pi.lean index 656f1dfa06252..916174981c99f 100644 --- a/src/data/multiset/pi.lean +++ b/src/data/multiset/pi.lean @@ -50,6 +50,25 @@ begin all_goals { simp [*, pi.cons_same, pi.cons_ne] }, end +@[simp] +lemma pi.cons_eta {m : multiset α} {a : α} (f : Π a' ∈ a ::ₘ m, δ a') : + pi.cons m a (f _ (mem_cons_self _ _)) (λ a' ha', f a' (mem_cons_of_mem ha')) = f := +begin + ext a' h', + by_cases a' = a, + { subst h, rw [pi.cons_same] }, + { rw [pi.cons_ne _ h] } +end + +lemma pi.cons_injective {a : α} {b : δ a} {s : multiset α} (hs : a ∉ s) : + function.injective (pi.cons s a b) := +assume f₁ f₂ eq, funext $ assume a', funext $ assume h', +have ne : a ≠ a', from assume h, hs $ h.symm ▸ h', +have a' ∈ a ::ₘ s, from mem_cons_of_mem h', +calc f₁ a' h' = pi.cons s a b f₁ a' this : by rw [pi.cons_ne this ne.symm] + ... = pi.cons s a b f₂ a' this : by rw [eq] + ... = f₂ a' h' : by rw [pi.cons_ne this ne.symm] + /-- `pi m t` constructs the Cartesian product over `t` indexed by `m`. -/ def pi (m : multiset α) (t : Πa, multiset (β a)) : multiset (Πa∈m, β a) := m.rec_on {pi.empty β} (λa m (p : multiset (Πa∈m, β a)), (t a).bind $ λb, p.map $ pi.cons m a b) @@ -73,15 +92,6 @@ end pi (a ::ₘ m) t = ((t a).bind $ λb, (pi m t).map $ pi.cons m a b) := rec_on_cons a m -lemma pi_cons_injective {a : α} {b : δ a} {s : multiset α} (hs : a ∉ s) : - function.injective (pi.cons s a b) := -assume f₁ f₂ eq, funext $ assume a', funext $ assume h', -have ne : a ≠ a', from assume h, hs $ h.symm ▸ h', -have a' ∈ a ::ₘ s, from mem_cons_of_mem h', -calc f₁ a' h' = pi.cons s a b f₁ a' this : by rw [pi.cons_ne this ne.symm] - ... = pi.cons s a b f₂ a' this : by rw [eq] - ... = f₂ a' h' : by rw [pi.cons_ne this ne.symm] - lemma card_pi (m : multiset α) (t : Πa, multiset (β a)) : card (pi m t) = prod (m.map $ λa, card (t a)) := multiset.induction_on m (by simp) (by simp [mul_comm] {contextual := tt}) @@ -94,7 +104,7 @@ begin have has : a ∉ s, by simp at hs; exact hs.1, have hs : nodup s, by simp at hs; exact hs.2, simp, - refine ⟨λ b hb, (ih hs $ λ a' h', ht a' $ mem_cons_of_mem h').map (pi_cons_injective has), _⟩, + refine ⟨λ b hb, (ih hs $ λ a' h', ht a' $ mem_cons_of_mem h').map (pi.cons_injective has), _⟩, refine (ht a $ mem_cons_self _ _).pairwise _, from assume b₁ hb₁ b₂ hb₂ neb, disjoint_map_map.2 (assume f hf g hg eq, have pi.cons s a b₁ f a (mem_cons_self _ _) = pi.cons s a b₂ g a (mem_cons_self _ _), @@ -102,16 +112,6 @@ begin neb $ show b₁ = b₂, by rwa [pi.cons_same, pi.cons_same] at this) end -@[simp] -lemma pi.cons_ext {m : multiset α} {a : α} (f : Π a' ∈ a ::ₘ m, δ a') : - pi.cons m a (f _ (mem_cons_self _ _)) (λ a' ha', f a' (mem_cons_of_mem ha')) = f := -begin - ext a' h', - by_cases a' = a, - { subst h, rw [pi.cons_same] }, - { rw [pi.cons_ne _ h] } -end - lemma mem_pi (m : multiset α) (t : Πa, multiset (β a)) : ∀f:Πa∈m, β a, (f ∈ pi m t) ↔ (∀a (h : a ∈ m), f a h ∈ t a) := begin @@ -126,7 +126,7 @@ begin { rw [pi.cons_ne _ h], apply hf' } }, { intro hf, refine ⟨_, hf a (mem_cons_self _ _), _, λ a ha, hf a (mem_cons_of_mem ha), _⟩, - rw pi.cons_ext } + rw pi.cons_eta } end end pi From 1d4d3ca5ec44693640c4f5e407a6b611f77accc8 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sun, 21 May 2023 22:57:58 +0000 Subject: [PATCH 1049/1291] =?UTF-8?q?feat(combinatorics/simple=5Fgraph):?= =?UTF-8?q?=20Szemer=C3=A9di's=20Regularity=20Lemma=20(#11000)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Prove (finally!) the equitable version of the Szemerédi Regularity Lemma. Co-authored-by: Bhavik Mehta --- .../simple_graph/regularity/lemma.lean | 152 ++++++++++++++++++ 1 file changed, 152 insertions(+) create mode 100644 src/combinatorics/simple_graph/regularity/lemma.lean diff --git a/src/combinatorics/simple_graph/regularity/lemma.lean b/src/combinatorics/simple_graph/regularity/lemma.lean new file mode 100644 index 0000000000000..0b271479520e1 --- /dev/null +++ b/src/combinatorics/simple_graph/regularity/lemma.lean @@ -0,0 +1,152 @@ +/- +Copyright (c) 2021 Yaël Dillies, Bhavik Mehta. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies, Bhavik Mehta +-/ +import combinatorics.simple_graph.regularity.increment + +/-! +# Szemerédi's Regularity Lemma + +In this file, we prove Szemerédi's Regularity Lemma (aka SRL). This is a landmark result in +combinatorics roughly stating that any sufficiently big graph behaves like a random graph. This is +useful because random graphs are well-behaved in many aspects. + +More precisely, SRL states that for any `ε > 0` and integer `l` there exists a bound `M` such that +any graph on at least `l` vertices can be partitioned into at least `l` parts and at most `M` parts +such that the resulting partitioned graph is `ε`-uniform. + +This statement is very robust to tweaking and many different versions exist. Here, we prove the +version where the resulting partition is equitable (aka an *equipartition*), namely all parts have +the same size up to a difference of `1`. + +The proof we formalise goes as follows: +1. Define an auxiliary measure of edge density, the *energy* of a partition. +2. Start with an arbitrary equipartition of size `l`. +3. Repeatedly break up the parts of the current equipartition in a big but controlled number of + parts. The key point is to break along the witnesses of non-uniformity, so that a lesser portion + of the pairs of parts are non-`ε`-uniform. +4. Check that this results in an equipartition with an energy greater than the energy of the current + partition, plus some constant. +5. Since the energy is between zero and one, we can't run this process forever. Check that when the + process stops we have an `ε`-uniform equipartition. + +This file only contains the final result. The supporting material is spread across the +`combinatorics.simple_graph.regularity` folder: +* `combinatorics.simple_graph.regularity.bound`: Definition of the bound on the number of parts. + Numerical inequalities involving the lemma constants. +* `combinatorics.simple_graph.regularity.energy`: Definition of the energy of a simple graph along a + partition. +* `combinatorics.simple_graph.regularity.uniform`: Definition of uniformity of a simple graph along + a pair of parts and along a partition. +* `combinatorics.simple_graph.regularity.equitabilise`: Construction of an equipartition with + a prescribed number of parts of each size and almost refining a given partition. +* `combinatorics.simple_graph.regularity.chunk`: Break up one part of the current equipartition. + Check that density between non-uniform parts increases, and that density between uniform parts + doesn't decrease too much. +* `combinatorics.simple_graph.regularity.increment`: Gather all those broken up parts into the new + equipartition (aka *increment partition*). Check that energy increases by at least a fixed amount. +* `combinatorics.simple_graph.regularity.lemma`: Wrap everything up into an induction on the energy. + +## TODO + +We currently only prove the equipartition version of SRL. + +* Prove the diagonal version. +* Prove the degree version. +* Define the regularity of a partition and prove the corresponding version. + +## References + +[Yaël Dillies, Bhavik Mehta, *Formalising Szemerédi’s Regularity Lemma in Lean*][srl_itp] +-/ + +open finpartition finset fintype function szemeredi_regularity +open_locale classical + +variables {α : Type*} [fintype α] (G : simple_graph α) {ε : ℝ} {l : ℕ} + +/-- Effective **Szemerédi Regularity Lemma**: For any sufficiently large graph, there is an +`ε`-uniform equipartition of bounded size (where the bound does not depend on the graph). -/ +theorem szemeredi_regularity (hε : 0 < ε) (hl : l ≤ card α) : + ∃ P : finpartition univ, + P.is_equipartition ∧ l ≤ P.parts.card ∧ P.parts.card ≤ bound ε l ∧ P.is_uniform G ε := +begin + obtain hα | hα := le_total (card α) (bound ε l), + -- If `card α ≤ bound ε l`, then the partition into singletons is acceptable. + { refine ⟨⊥, bot_is_equipartition _, _⟩, + rw [card_bot, card_univ], + exact ⟨hl, hα, bot_is_uniform _ hε⟩ }, + -- Else, let's start from a dummy equipartition of size `initial_bound ε l`. + let t := initial_bound ε l, + have htα : t ≤ (univ : finset α).card := + (initial_bound_le_bound _ _).trans (by rwa finset.card_univ), + obtain ⟨dum, hdum₁, hdum₂⟩ := exists_equipartition_card_eq (univ : finset α) + (initial_bound_pos _ _).ne' htα, + obtain hε₁ | hε₁ := le_total 1 ε, + ---If `ε ≥ 1`, then this dummy equipartition is `ε`-uniform, so we're done. + { exact ⟨dum, hdum₁, (le_initial_bound ε l).trans hdum₂.ge, + hdum₂.le.trans (initial_bound_le_bound ε l), (dum.is_uniform_one G).mono hε₁⟩ }, + -- Else, set up the induction on energy. We phrase it through the existence for each `i` of an + -- equipartition of size bounded by `step_bound^[i]) (initial_bound ε l)` and which is either + -- `ε`-uniform or has energy at least `ε ^ 5 / 4 * i`. + haveI : nonempty α, + { rw ←fintype.card_pos_iff, + exact (bound_pos _ _).trans_le hα }, + suffices h : ∀ i, ∃ P : finpartition (univ : finset α), P.is_equipartition ∧ + t ≤ P.parts.card ∧ P.parts.card ≤ (step_bound^[i]) t ∧ + (P.is_uniform G ε ∨ ε ^ 5 / 4 * i ≤ P.energy G), + -- For `i > 4 / ε ^ 5` we know that the partition we get can't have energy `≥ ε ^ 5 / 4 * i > 1`, + -- so it must instead be `ε`-uniform and we won. + { obtain ⟨P, hP₁, hP₂, hP₃, hP₄⟩ := h (⌊4 / ε ^ 5⌋₊ + 1), + refine ⟨P, hP₁, (le_initial_bound _ _).trans hP₂, hP₃.trans _, + hP₄.resolve_right $ λ hPenergy, lt_irrefl (1 : ℝ) _⟩, + { rw iterate_succ_apply', + exact mul_le_mul_left' (pow_le_pow_of_le_left (by norm_num) (by norm_num) _) _ }, + calc + 1 = ε ^ 5 / 4 * (4 / ε ^ 5) + : by { rw [mul_comm, div_mul_div_cancel 4 (pow_pos hε 5).ne'], norm_num } + ... < ε ^ 5 / 4 * (⌊4 / ε ^ 5⌋₊ + 1) + : (mul_lt_mul_left $ by positivity).2 (nat.lt_floor_add_one _) + ... ≤ (P.energy G : ℝ) : by rwa ←nat.cast_add_one + ... ≤ 1 : by exact_mod_cast P.energy_le_one G }, + -- Let's do the actual induction. + intro i, + induction i with i ih, + -- For `i = 0`, the dummy equipartition is enough. + { refine ⟨dum, hdum₁, hdum₂.ge, hdum₂.le, or.inr _⟩, + rw [nat.cast_zero, mul_zero], + exact_mod_cast dum.energy_nonneg G }, + -- For the induction step at `i + 1`, find `P` the equipartition at `i`. + obtain ⟨P, hP₁, hP₂, hP₃, hP₄⟩ := ih, + by_cases huniform : P.is_uniform G ε, + -- If `P` is already uniform, then no need to break it up further. We can just return `P` again. + { refine ⟨P, hP₁, hP₂, _, or.inl huniform⟩, + rw iterate_succ_apply', + exact hP₃.trans (le_step_bound _) }, + -- Else, `P` must instead have energy at least `ε ^ 5 / 4 * i`. + replace hP₄ := hP₄.resolve_left huniform, + -- We gather a few numerical facts. + have hεl' : 100 < 4 ^ P.parts.card * ε ^ 5, + { exact (hundred_lt_pow_initial_bound_mul hε l).trans_le + (mul_le_mul_of_nonneg_right (pow_le_pow (by norm_num) hP₂) $ by positivity) }, + have hi : (i : ℝ) ≤ 4 / ε ^ 5, + { have hi : ε ^ 5 / 4 * ↑i ≤ 1 := hP₄.trans (by exact_mod_cast P.energy_le_one G), + rw [div_mul_eq_mul_div, div_le_iff (show (0:ℝ) < 4, by norm_num)] at hi, + norm_num at hi, + rwa le_div_iff' (pow_pos hε _) }, + have hsize : P.parts.card ≤ (step_bound^[⌊4 / ε ^ 5⌋₊] t) := + hP₃.trans (monotone_iterate_of_id_le le_step_bound (nat.le_floor hi) _), + have hPα : P.parts.card * 16 ^ P.parts.card ≤ card α := + (nat.mul_le_mul hsize (nat.pow_le_pow_of_le_right (by norm_num) hsize)).trans hα, + -- We return the increment equipartition of `P`, which has energy `≥ ε ^ 5 / 4 * (i + 1)`. + refine ⟨increment hP₁ G ε, increment_is_equipartition hP₁ G ε, _, _, + or.inr $ le_trans _ $ energy_increment hP₁ ((seven_le_initial_bound ε l).trans hP₂) + hεl' hPα huniform hε₁⟩, + { rw card_increment hPα huniform, + exact hP₂.trans (le_step_bound _) }, + { rw [card_increment hPα huniform, iterate_succ_apply'], + exact step_bound_mono hP₃ }, + { rw [nat.cast_succ, mul_add, mul_one], + exact add_le_add_right hP₄ _ } +end From 4280f5f32e16755ec7985ce11e189b6cd6ff6735 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Mon, 22 May 2023 05:47:04 +0000 Subject: [PATCH 1050/1291] chore(*): add mathlib4 synchronization comments (#19057) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.category.Group.limits` * `algebra.category.Module.kernels` * `algebra.category.Mon.filtered_colimits` * `algebra.char_p.mixed_char_zero` * `algebraic_geometry.presheafed_space` * `algebraic_geometry.projective_spectrum.topology` * `analysis.calculus.fderiv.basic` * `analysis.normed_space.star.mul` * `analysis.special_functions.complex.circle` * `analysis.special_functions.pow.nnreal` * `linear_algebra.matrix.charpoly.basic` * `linear_algebra.matrix.charpoly.coeff` * `linear_algebra.matrix.charpoly.linear_map` * `measure_theory.constructions.borel_space.basic` * `measure_theory.constructions.borel_space.complex` * `measure_theory.constructions.borel_space.continuous_linear_map` * `measure_theory.constructions.polish` * `measure_theory.function.ae_measurable_order` * `measure_theory.function.ess_sup` * `measure_theory.function.floor` * `measure_theory.function.simple_func` * `measure_theory.function.simple_func_dense` * `measure_theory.function.special_functions.basic` * `measure_theory.integral.lebesgue` * `measure_theory.measure.regular` * `measure_theory.measure.stieltjes` * `number_theory.class_number.admissible_card_pow_degree` * `ring_theory.graded_algebra.homogeneous_ideal` * `ring_theory.graded_algebra.homogeneous_localization` * `ring_theory.ideal.minimal_prime` * `ring_theory.polynomial_algebra` * `topology.algebra.module.locally_convex` * `topology.continuous_function.bounded` * `topology.sheaves.sheaf_condition.sites` --- src/algebra/category/Group/limits.lean | 3 +++ src/algebra/category/Module/kernels.lean | 3 +++ src/algebra/category/Mon/filtered_colimits.lean | 3 +++ src/algebra/char_p/mixed_char_zero.lean | 3 +++ src/algebraic_geometry/presheafed_space.lean | 3 +++ src/algebraic_geometry/projective_spectrum/topology.lean | 3 +++ src/analysis/calculus/fderiv/basic.lean | 3 +++ src/analysis/normed_space/star/mul.lean | 5 ++++- src/analysis/special_functions/complex/circle.lean | 3 +++ src/analysis/special_functions/pow/nnreal.lean | 3 +++ src/linear_algebra/matrix/charpoly/basic.lean | 3 +++ src/linear_algebra/matrix/charpoly/coeff.lean | 3 +++ src/linear_algebra/matrix/charpoly/linear_map.lean | 3 +++ src/measure_theory/constructions/borel_space/basic.lean | 3 +++ src/measure_theory/constructions/borel_space/complex.lean | 5 ++++- .../constructions/borel_space/continuous_linear_map.lean | 3 +++ src/measure_theory/constructions/polish.lean | 3 +++ src/measure_theory/function/ae_measurable_order.lean | 3 +++ src/measure_theory/function/ess_sup.lean | 3 +++ src/measure_theory/function/floor.lean | 3 +++ src/measure_theory/function/simple_func.lean | 3 +++ src/measure_theory/function/simple_func_dense.lean | 3 +++ src/measure_theory/function/special_functions/basic.lean | 3 +++ src/measure_theory/integral/lebesgue.lean | 3 +++ src/measure_theory/measure/regular.lean | 3 +++ src/measure_theory/measure/stieltjes.lean | 3 +++ .../class_number/admissible_card_pow_degree.lean | 3 +++ src/ring_theory/graded_algebra/homogeneous_ideal.lean | 3 +++ src/ring_theory/graded_algebra/homogeneous_localization.lean | 3 +++ src/ring_theory/ideal/minimal_prime.lean | 3 +++ src/ring_theory/polynomial_algebra.lean | 3 +++ src/topology/algebra/module/locally_convex.lean | 3 +++ src/topology/continuous_function/bounded.lean | 3 +++ src/topology/sheaves/sheaf_condition/sites.lean | 3 +++ 34 files changed, 104 insertions(+), 2 deletions(-) diff --git a/src/algebra/category/Group/limits.lean b/src/algebra/category/Group/limits.lean index 7b726cfc1cd1b..dfd5969dd3b2d 100644 --- a/src/algebra/category/Group/limits.lean +++ b/src/algebra/category/Group/limits.lean @@ -12,6 +12,9 @@ import category_theory.concrete_category.elementwise /-! # The category of (commutative) (additive) groups has all limits +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Further, these limits are preserved by the forgetful functor --- that is, the underlying types are just the limits in the category of types. diff --git a/src/algebra/category/Module/kernels.lean b/src/algebra/category/Module/kernels.lean index 1b53ef3cd63fd..80d2b4986e2f7 100644 --- a/src/algebra/category/Module/kernels.lean +++ b/src/algebra/category/Module/kernels.lean @@ -8,6 +8,9 @@ import category_theory.concrete_category.elementwise /-! # The concrete (co)kernels in the category of modules are (co)kernels in the categorical sense. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ open category_theory diff --git a/src/algebra/category/Mon/filtered_colimits.lean b/src/algebra/category/Mon/filtered_colimits.lean index 37fd60e8dcbda..93d263d8b0941 100644 --- a/src/algebra/category/Mon/filtered_colimits.lean +++ b/src/algebra/category/Mon/filtered_colimits.lean @@ -11,6 +11,9 @@ import category_theory.limits.types /-! # The forgetful functor from (commutative) (additive) monoids preserves filtered colimits. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Forgetful functors from algebraic categories usually don't preserve colimits. However, they tend to preserve _filtered_ colimits. diff --git a/src/algebra/char_p/mixed_char_zero.lean b/src/algebra/char_p/mixed_char_zero.lean index eb084b3e92e85..0c17a7d83dc45 100644 --- a/src/algebra/char_p/mixed_char_zero.lean +++ b/src/algebra/char_p/mixed_char_zero.lean @@ -11,6 +11,9 @@ import tactic.field_simp /-! # Equal and mixed characteristic +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In commutative algebra, some statments are simpler when working over a `ℚ`-algebra `R`, in which case one also says that the ring has "equal characteristic zero". A ring that is not a `ℚ`-algebra has either positive characteristic or there exists a prime ideal `I ⊂ R` such that diff --git a/src/algebraic_geometry/presheafed_space.lean b/src/algebraic_geometry/presheafed_space.lean index 575dca4468f38..9c38f1afa9515 100644 --- a/src/algebraic_geometry/presheafed_space.lean +++ b/src/algebraic_geometry/presheafed_space.lean @@ -9,6 +9,9 @@ import category_theory.adjunction.fully_faithful /-! # Presheafed spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Introduces the category of topological spaces equipped with a presheaf (taking values in an arbitrary target category `C`.) diff --git a/src/algebraic_geometry/projective_spectrum/topology.lean b/src/algebraic_geometry/projective_spectrum/topology.lean index cd87b1507213d..d6d8f3f9408c0 100644 --- a/src/algebraic_geometry/projective_spectrum/topology.lean +++ b/src/algebraic_geometry/projective_spectrum/topology.lean @@ -11,6 +11,9 @@ import topology.sets.opens /-! # Projective spectrum of a graded ring +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The projective spectrum of a graded commutative ring is the subtype of all homogenous ideals that are prime and do not contain the irrelevant ideal. It is naturally endowed with a topology: the Zariski topology. diff --git a/src/analysis/calculus/fderiv/basic.lean b/src/analysis/calculus/fderiv/basic.lean index f97d738340d5f..566a26ac49a15 100644 --- a/src/analysis/calculus/fderiv/basic.lean +++ b/src/analysis/calculus/fderiv/basic.lean @@ -10,6 +10,9 @@ import analysis.normed_space.bounded_linear_maps /-! # The Fréchet derivative +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Let `E` and `F` be normed spaces, `f : E → F`, and `f' : E →L[𝕜] F` a continuous 𝕜-linear map, where `𝕜` is a non-discrete normed field. Then diff --git a/src/analysis/normed_space/star/mul.lean b/src/analysis/normed_space/star/mul.lean index 46e1f84950911..bdf88d0d85bee 100644 --- a/src/analysis/normed_space/star/mul.lean +++ b/src/analysis/normed_space/star/mul.lean @@ -6,7 +6,10 @@ Authors: Jireh Loreaux import analysis.normed_space.star.basic import analysis.normed_space.operator_norm -/-! # The left-regular representation is an isometry for C⋆-algebras -/ +/-! # The left-regular representation is an isometry for C⋆-algebras + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4.-/ open continuous_linear_map diff --git a/src/analysis/special_functions/complex/circle.lean b/src/analysis/special_functions/complex/circle.lean index 69c912ada3427..9da2facb23e57 100644 --- a/src/analysis/special_functions/complex/circle.lean +++ b/src/analysis/special_functions/complex/circle.lean @@ -9,6 +9,9 @@ import analysis.special_functions.complex.log /-! # Maps on the unit circle +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove some basic lemmas about `exp_map_circle` and the restriction of `complex.arg` to the unit circle. These two maps define a local equivalence between `circle` and `ℝ`, see `circle.arg_local_equiv` and `circle.arg_equiv`, that sends the whole circle to `(-π, π]`. diff --git a/src/analysis/special_functions/pow/nnreal.lean b/src/analysis/special_functions/pow/nnreal.lean index 65aec15568256..680cad3232806 100644 --- a/src/analysis/special_functions/pow/nnreal.lean +++ b/src/analysis/special_functions/pow/nnreal.lean @@ -9,6 +9,9 @@ import analysis.special_functions.pow.real /-! # Power function on `ℝ≥0` and `ℝ≥0∞` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We construct the power functions `x ^ y` where * `x` is a nonnegative real number and `y` is a real number; * `x` is a number from `[0, +∞]` (a.k.a. `ℝ≥0∞`) and `y` is a real number. diff --git a/src/linear_algebra/matrix/charpoly/basic.lean b/src/linear_algebra/matrix/charpoly/basic.lean index 36ab7df851783..f402becdd93d6 100644 --- a/src/linear_algebra/matrix/charpoly/basic.lean +++ b/src/linear_algebra/matrix/charpoly/basic.lean @@ -11,6 +11,9 @@ import tactic.squeeze /-! # Characteristic polynomials and the Cayley-Hamilton theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define characteristic polynomials of matrices and prove the Cayley–Hamilton theorem over arbitrary commutative rings. diff --git a/src/linear_algebra/matrix/charpoly/coeff.lean b/src/linear_algebra/matrix/charpoly/coeff.lean index f92caae539681..834300d00a30b 100644 --- a/src/linear_algebra/matrix/charpoly/coeff.lean +++ b/src/linear_algebra/matrix/charpoly/coeff.lean @@ -10,6 +10,9 @@ import linear_algebra.matrix.charpoly.basic /-! # Characteristic polynomials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We give methods for computing coefficients of the characteristic polynomial. ## Main definitions diff --git a/src/linear_algebra/matrix/charpoly/linear_map.lean b/src/linear_algebra/matrix/charpoly/linear_map.lean index ca434773697f6..d37988acc241b 100644 --- a/src/linear_algebra/matrix/charpoly/linear_map.lean +++ b/src/linear_algebra/matrix/charpoly/linear_map.lean @@ -10,6 +10,9 @@ import linear_algebra.matrix.to_lin # Calyley-Hamilton theorem for f.g. modules. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given a fixed finite spanning set `b : ι → M` of a `R`-module `M`, we say that a matrix `M` represents an endomorphism `f : M →ₗ[R] M` if the matrix as an endomorphism of `ι → R` commutes with `f` via the projection `(ι → R) →ₗ[R] M` given by `b`. diff --git a/src/measure_theory/constructions/borel_space/basic.lean b/src/measure_theory/constructions/borel_space/basic.lean index 9ca71671a64b3..73cbb8b06bdc0 100644 --- a/src/measure_theory/constructions/borel_space/basic.lean +++ b/src/measure_theory/constructions/borel_space/basic.lean @@ -19,6 +19,9 @@ import topology.semicontinuous /-! # Borel (measurable) space +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions * `borel α` : the least `σ`-algebra that contains all open sets; diff --git a/src/measure_theory/constructions/borel_space/complex.lean b/src/measure_theory/constructions/borel_space/complex.lean index 1a69089f7ab98..de7034d99a042 100644 --- a/src/measure_theory/constructions/borel_space/complex.lean +++ b/src/measure_theory/constructions/borel_space/complex.lean @@ -6,7 +6,10 @@ Authors: Yury Kudryashov import analysis.complex.basic import measure_theory.constructions.borel_space.basic -/-! # Equip `ℂ` with the Borel sigma-algebra -/ +/-! # Equip `ℂ` with the Borel sigma-algebra + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4.-/ noncomputable theory diff --git a/src/measure_theory/constructions/borel_space/continuous_linear_map.lean b/src/measure_theory/constructions/borel_space/continuous_linear_map.lean index 6d7908b94a05a..ee21c514a2f3b 100644 --- a/src/measure_theory/constructions/borel_space/continuous_linear_map.lean +++ b/src/measure_theory/constructions/borel_space/continuous_linear_map.lean @@ -9,6 +9,9 @@ import measure_theory.constructions.borel_space.basic /-! # Measurable functions in normed spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + -/ open measure_theory diff --git a/src/measure_theory/constructions/polish.lean b/src/measure_theory/constructions/polish.lean index 9c7d3638995c9..720a7aac73281 100644 --- a/src/measure_theory/constructions/polish.lean +++ b/src/measure_theory/constructions/polish.lean @@ -10,6 +10,9 @@ import measure_theory.constructions.borel_space.basic /-! # The Borel sigma-algebra on Polish spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We discuss several results pertaining to the relationship between the topology and the Borel structure on Polish spaces. diff --git a/src/measure_theory/function/ae_measurable_order.lean b/src/measure_theory/function/ae_measurable_order.lean index 478ef9275ea72..fd40f966cee90 100644 --- a/src/measure_theory/function/ae_measurable_order.lean +++ b/src/measure_theory/function/ae_measurable_order.lean @@ -8,6 +8,9 @@ import measure_theory.constructions.borel_space.basic /-! # Measurability criterion for ennreal-valued functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Consider a function `f : α → ℝ≥0∞`. If the level sets `{f < p}` and `{q < f}` have measurable supersets which are disjoint up to measure zero when `p` and `q` are finite numbers satisfying `p < q`, then `f` is almost-everywhere measurable. This is proved in diff --git a/src/measure_theory/function/ess_sup.lean b/src/measure_theory/function/ess_sup.lean index 01c891dcc97ea..101cfc2794317 100644 --- a/src/measure_theory/function/ess_sup.lean +++ b/src/measure_theory/function/ess_sup.lean @@ -8,6 +8,9 @@ import order.filter.ennreal /-! # Essential supremum and infimum + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. We define the essential supremum and infimum of a function `f : α → β` with respect to a measure `μ` on `α`. The essential supremum is the infimum of the constants `c : β` such that `f x ≤ c` almost everywhere. diff --git a/src/measure_theory/function/floor.lean b/src/measure_theory/function/floor.lean index 41b72efe283a2..1f7c974ad0c6e 100644 --- a/src/measure_theory/function/floor.lean +++ b/src/measure_theory/function/floor.lean @@ -8,6 +8,9 @@ import measure_theory.constructions.borel_space.basic /-! # Measurability of `⌊x⌋` etc +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove that `int.floor`, `int.ceil`, `int.fract`, `nat.floor`, and `nat.ceil` are measurable under some assumptions on the (semi)ring. -/ diff --git a/src/measure_theory/function/simple_func.lean b/src/measure_theory/function/simple_func.lean index 0c8816199c41d..7e69629179f8c 100644 --- a/src/measure_theory/function/simple_func.lean +++ b/src/measure_theory/function/simple_func.lean @@ -10,6 +10,9 @@ import algebra.support /-! # Simple functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A function `f` from a measurable space to any type is called *simple*, if every preimage `f ⁻¹' {x}` is measurable, and the range is finite. In this file, we define simple functions and establish their basic properties; and we construct a sequence of simple functions approximating an arbitrary Borel diff --git a/src/measure_theory/function/simple_func_dense.lean b/src/measure_theory/function/simple_func_dense.lean index b9d7038805441..25fcda5f6fa32 100644 --- a/src/measure_theory/function/simple_func_dense.lean +++ b/src/measure_theory/function/simple_func_dense.lean @@ -8,6 +8,9 @@ import measure_theory.function.simple_func /-! # Density of simple functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Show that each Borel measurable function can be approximated pointwise by a sequence of simple functions. diff --git a/src/measure_theory/function/special_functions/basic.lean b/src/measure_theory/function/special_functions/basic.lean index f2a4b2b4397e5..17665d0e7e0cc 100644 --- a/src/measure_theory/function/special_functions/basic.lean +++ b/src/measure_theory/function/special_functions/basic.lean @@ -10,6 +10,9 @@ import measure_theory.constructions.borel_space.complex /-! # Measurability of real and complex functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We show that most standard real and complex functions are measurable, notably `exp`, `cos`, `sin`, `cosh`, `sinh`, `log`, `pow`, `arcsin`, `arccos`. diff --git a/src/measure_theory/integral/lebesgue.lean b/src/measure_theory/integral/lebesgue.lean index 413dc0a066358..0af668bac16f0 100644 --- a/src/measure_theory/integral/lebesgue.lean +++ b/src/measure_theory/integral/lebesgue.lean @@ -10,6 +10,9 @@ import measure_theory.measure.mutually_singular /-! # Lower Lebesgue integral for `ℝ≥0∞`-valued functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define the lower Lebesgue integral of an `ℝ≥0∞`-valued function. ## Notation diff --git a/src/measure_theory/measure/regular.lean b/src/measure_theory/measure/regular.lean index 4f3eb55b5584f..26bb77f77822e 100644 --- a/src/measure_theory/measure/regular.lean +++ b/src/measure_theory/measure/regular.lean @@ -9,6 +9,9 @@ import measure_theory.constructions.borel_space.basic /-! # Regular measures +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A measure is `outer_regular` if the measure of any measurable set `A` is the infimum of `μ U` over all open sets `U` containing `A`. diff --git a/src/measure_theory/measure/stieltjes.lean b/src/measure_theory/measure/stieltjes.lean index 7115704cf6288..cd65059c9a0bc 100644 --- a/src/measure_theory/measure/stieltjes.lean +++ b/src/measure_theory/measure/stieltjes.lean @@ -9,6 +9,9 @@ import topology.algebra.order.left_right_lim /-! # Stieltjes measures on the real line +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Consider a function `f : ℝ → ℝ` which is monotone and right-continuous. Then one can define a corrresponding measure, giving mass `f b - f a` to the interval `(a, b]`. diff --git a/src/number_theory/class_number/admissible_card_pow_degree.lean b/src/number_theory/class_number/admissible_card_pow_degree.lean index 37a7b22bbc3d3..dfb55c2f566e6 100644 --- a/src/number_theory/class_number/admissible_card_pow_degree.lean +++ b/src/number_theory/class_number/admissible_card_pow_degree.lean @@ -11,6 +11,9 @@ import data.polynomial.degree.card_pow_degree /-! # Admissible absolute values on polynomials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines an admissible absolute value `polynomial.card_pow_degree_is_admissible` which we use to show the class number of the ring of integers of a function field is finite. diff --git a/src/ring_theory/graded_algebra/homogeneous_ideal.lean b/src/ring_theory/graded_algebra/homogeneous_ideal.lean index 02156a1661fa2..535d934b73de8 100644 --- a/src/ring_theory/graded_algebra/homogeneous_ideal.lean +++ b/src/ring_theory/graded_algebra/homogeneous_ideal.lean @@ -10,6 +10,9 @@ import ring_theory.graded_algebra.basic /-! # Homogeneous ideals of a graded algebra +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines homogeneous ideals of `graded_ring 𝒜` where `𝒜 : ι → submodule R A` and operations on them. diff --git a/src/ring_theory/graded_algebra/homogeneous_localization.lean b/src/ring_theory/graded_algebra/homogeneous_localization.lean index f311442ab436f..0fcacd0ccbdc9 100644 --- a/src/ring_theory/graded_algebra/homogeneous_localization.lean +++ b/src/ring_theory/graded_algebra/homogeneous_localization.lean @@ -9,6 +9,9 @@ import ring_theory.graded_algebra.basic /-! # Homogeneous Localization +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Notation - `ι` is a commutative monoid; - `R` is a commutative semiring; diff --git a/src/ring_theory/ideal/minimal_prime.lean b/src/ring_theory/ideal/minimal_prime.lean index 785a283f8026a..bcdf0bd853eb7 100644 --- a/src/ring_theory/ideal/minimal_prime.lean +++ b/src/ring_theory/ideal/minimal_prime.lean @@ -11,6 +11,9 @@ import order.minimal # Minimal primes +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We provide various results concerning the minimal primes above an ideal ## Main results diff --git a/src/ring_theory/polynomial_algebra.lean b/src/ring_theory/polynomial_algebra.lean index c77c47dce0826..7e08efc48eed2 100644 --- a/src/ring_theory/polynomial_algebra.lean +++ b/src/ring_theory/polynomial_algebra.lean @@ -11,6 +11,9 @@ import data.matrix.dmatrix /-! # Algebra isomorphism between matrices of polynomials and polynomials of matrices +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given `[comm_ring R] [ring A] [algebra R A]` we show `A[X] ≃ₐ[R] (A ⊗[R] R[X])`. Combining this with the isomorphism `matrix n n A ≃ₐ[R] (A ⊗[R] matrix n n R)` proved earlier diff --git a/src/topology/algebra/module/locally_convex.lean b/src/topology/algebra/module/locally_convex.lean index 11ea2c947006b..4536bd26514be 100644 --- a/src/topology/algebra/module/locally_convex.lean +++ b/src/topology/algebra/module/locally_convex.lean @@ -7,6 +7,9 @@ import analysis.convex.topology /-! # Locally convex topological modules +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A `locally_convex_space` is a topological semimodule over an ordered semiring in which any point admits a neighborhood basis made of convex sets, or equivalently, in which convex neighborhoods of a point form a neighborhood basis at that point. diff --git a/src/topology/continuous_function/bounded.lean b/src/topology/continuous_function/bounded.lean index 7207cb70e936d..3563a0c5729c6 100644 --- a/src/topology/continuous_function/bounded.lean +++ b/src/topology/continuous_function/bounded.lean @@ -13,6 +13,9 @@ import topology.metric_space.equicontinuity /-! # Bounded continuous functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The type of bounded continuous functions taking values in a metric space, with the uniform distance. diff --git a/src/topology/sheaves/sheaf_condition/sites.lean b/src/topology/sheaves/sheaf_condition/sites.lean index 815f909b02492..d6692f330c357 100644 --- a/src/topology/sheaves/sheaf_condition/sites.lean +++ b/src/topology/sheaves/sheaf_condition/sites.lean @@ -12,6 +12,9 @@ import category_theory.sites.dense_subsite # Coverings and sieves; from sheaves on sites and sheaves on spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we connect coverings in a topological space to sieves in the associated Grothendieck topology, in preparation of connecting the sheaf condition on sites to the various sheaf conditions on spaces. From 13bce9a6b6c44f6b4c91ac1c1d2a816e2533d395 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 22 May 2023 08:06:10 +0000 Subject: [PATCH 1051/1291] chore(linear_algebra/basis): `simp` lemmas about `basis.equiv_fun` (#19021) --- src/analysis/inner_product_space/pi_L2.lean | 7 +------ src/analysis/normed_space/pi_Lp.lean | 4 ++++ src/linear_algebra/basis.lean | 7 +++++++ src/linear_algebra/std_basis.lean | 3 +++ 4 files changed, 15 insertions(+), 6 deletions(-) diff --git a/src/analysis/inner_product_space/pi_L2.lean b/src/analysis/inner_product_space/pi_L2.lean index 011794e4a6975..93730db0aff12 100644 --- a/src/analysis/inner_product_space/pi_L2.lean +++ b/src/analysis/inner_product_space/pi_L2.lean @@ -334,12 +334,7 @@ end @[simp] protected lemma coe_to_basis_repr (b : orthonormal_basis ι 𝕜 E) : b.to_basis.equiv_fun = b.repr.to_linear_equiv := -begin - change (basis.of_equiv_fun b.repr.to_linear_equiv).equiv_fun = b.repr.to_linear_equiv, - ext x j, - simp only [basis.of_equiv_fun_repr_apply, linear_isometry_equiv.coe_to_linear_equiv, - basis.equiv_fun_apply], -end +basis.equiv_fun_of_equiv_fun _ @[simp] protected lemma coe_to_basis_repr_apply (b : orthonormal_basis ι 𝕜 E) (x : E) (i : ι) : b.to_basis.repr x i = b.repr x i := diff --git a/src/analysis/normed_space/pi_Lp.lean b/src/analysis/normed_space/pi_Lp.lean index 05a6a28014559..dc735f9d255b2 100644 --- a/src/analysis/normed_space/pi_Lp.lean +++ b/src/analysis/normed_space/pi_Lp.lean @@ -793,6 +793,10 @@ by simp_rw [basis_fun, basis.coe_of_equiv_fun, pi_Lp.linear_equiv_symm_apply, pi (basis_fun p 𝕜 ι).repr x i = x i := rfl +@[simp] lemma basis_fun_equiv_fun : + (basis_fun p 𝕜 ι).equiv_fun = pi_Lp.linear_equiv p 𝕜 (λ _ : ι, 𝕜) := +basis.equiv_fun_of_equiv_fun _ + lemma basis_fun_eq_pi_basis_fun : basis_fun p 𝕜 ι = (pi.basis_fun 𝕜 ι).map (pi_Lp.linear_equiv p 𝕜 (λ _ : ι, 𝕜)).symm := rfl diff --git a/src/linear_algebra/basis.lean b/src/linear_algebra/basis.lean index c560d1cfc33ec..c7f2f38be3cd7 100644 --- a/src/linear_algebra/basis.lean +++ b/src/linear_algebra/basis.lean @@ -833,6 +833,13 @@ begin simp only [finset.mem_univ, if_true, pi.zero_apply, one_smul, finset.sum_ite_eq', zero_smul], end +@[simp] lemma basis.equiv_fun_of_equiv_fun (e : M ≃ₗ[R] (ι → R)) : + (basis.of_equiv_fun e).equiv_fun = e := +begin + ext j, + simp_rw [basis.equiv_fun_apply, basis.of_equiv_fun_repr_apply], +end + variables (S : Type*) [semiring S] [module S M'] variables [smul_comm_class R S M'] diff --git a/src/linear_algebra/std_basis.lean b/src/linear_algebra/std_basis.lean index f1af076104f9f..56f2224deb91e 100644 --- a/src/linear_algebra/std_basis.lean +++ b/src/linear_algebra/std_basis.lean @@ -255,6 +255,9 @@ by { simp only [basis_fun, basis.coe_of_equiv_fun, linear_equiv.refl_symm, (pi.basis_fun R η).repr x i = x i := by simp [basis_fun] +@[simp] lemma basis_fun_equiv_fun : (pi.basis_fun R η).equiv_fun = linear_equiv.refl _ _ := +basis.equiv_fun_of_equiv_fun _ + end end module From 3b88f4005dc2e28d42f974cc1ce838f0dafb39b8 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Mon, 22 May 2023 10:23:18 +0000 Subject: [PATCH 1052/1291] feat(probability/kernel/disintegration): disintegration of finite measures on product spaces (#18834) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Disintegration of finite measures on `α × Ω`, where `Ω` is a standard Borel space. Co-authored-by: RemyDegenne --- .../constructions/prod/basic.lean | 29 ++ src/probability/kernel/cond_cdf.lean | 45 ++ src/probability/kernel/disintegration.lean | 411 ++++++++++++++++++ 3 files changed, 485 insertions(+) create mode 100644 src/probability/kernel/disintegration.lean diff --git a/src/measure_theory/constructions/prod/basic.lean b/src/measure_theory/constructions/prod/basic.lean index f2afcf79db87c..b4ef7bf51ed5b 100644 --- a/src/measure_theory/constructions/prod/basic.lean +++ b/src/measure_theory/constructions/prod/basic.lean @@ -189,6 +189,35 @@ begin exact measurable_measure_prod_mk_right hs end +lemma measurable_embedding.prod_mk {α β γ δ : Type*} {mα : measurable_space α} + {mβ : measurable_space β} {mγ : measurable_space γ} {mδ : measurable_space δ} + {f : α → β} {g : γ → δ} (hg : measurable_embedding g) (hf : measurable_embedding f) : + measurable_embedding (λ x : γ × α, (g x.1, f x.2)) := +begin + have h_inj : function.injective (λ x : γ × α, (g x.fst, f x.snd)), + { intros x y hxy, + rw [← @prod.mk.eta _ _ x, ← @prod.mk.eta _ _ y], + simp only [prod.mk.inj_iff] at hxy ⊢, + exact ⟨hg.injective hxy.1, hf.injective hxy.2⟩, }, + refine ⟨h_inj, _, _⟩, + { exact (hg.measurable.comp measurable_fst).prod_mk (hf.measurable.comp measurable_snd), }, + { -- Induction using the π-system of rectangles + refine λ s hs, @measurable_space.induction_on_inter _ + (λ s, measurable_set ((λ (x : γ × α), (g x.fst, f x.snd)) '' s)) _ _ generate_from_prod.symm + is_pi_system_prod _ _ _ _ _ hs, + { simp only [set.image_empty, measurable_set.empty], }, + { rintros t ⟨t₁, t₂, ht₁, ht₂, rfl⟩, + rw ← set.prod_image_image_eq, + exact (hg.measurable_set_image.mpr ht₁).prod (hf.measurable_set_image.mpr ht₂), }, + { intros t ht ht_m, + rw [← set.range_diff_image h_inj, ← set.prod_range_range_eq], + exact measurable_set.diff + (measurable_set.prod hg.measurable_set_range hf.measurable_set_range) ht_m, }, + { intros g hg_disj hg_meas hg, + simp_rw set.image_Union, + exact measurable_set.Union hg, }, }, +end + /-- The Lebesgue integral is measurable. This shows that the integrand of (the right-hand-side of) Tonelli's theorem is measurable. -/ lemma measurable.lintegral_prod_right' [sigma_finite ν] : diff --git a/src/probability/kernel/cond_cdf.lean b/src/probability/kernel/cond_cdf.lean index a98bbd76489a1..e9478de11a3c7 100644 --- a/src/probability/kernel/cond_cdf.lean +++ b/src/probability/kernel/cond_cdf.lean @@ -966,4 +966,49 @@ lemma integral_cond_cdf (ρ : measure (α × ℝ)) [is_finite_measure ρ] (x : ∫ a, cond_cdf ρ a x ∂ρ.fst = (ρ (univ ×ˢ Iic x)).to_real := by rw [← set_integral_cond_cdf ρ _ measurable_set.univ, measure.restrict_univ] +section measure + +lemma measure_cond_cdf_Iic (ρ : measure (α × ℝ)) (a : α) (x : ℝ) : + (cond_cdf ρ a).measure (Iic x) = ennreal.of_real (cond_cdf ρ a x) := +begin + rw [← sub_zero (cond_cdf ρ a x)], + exact (cond_cdf ρ a).measure_Iic (tendsto_cond_cdf_at_bot ρ a) _, +end + +lemma measure_cond_cdf_univ (ρ : measure (α × ℝ)) (a : α) : + (cond_cdf ρ a).measure univ = 1 := +begin + rw [← ennreal.of_real_one, ← sub_zero (1 : ℝ)], + exact stieltjes_function.measure_univ _ (tendsto_cond_cdf_at_bot ρ a) + (tendsto_cond_cdf_at_top ρ a), +end + +instance (ρ : measure (α × ℝ)) (a : α) : is_probability_measure ((cond_cdf ρ a).measure) := +⟨measure_cond_cdf_univ ρ a⟩ + +/-- The function `a ↦ (cond_cdf ρ a).measure` is measurable. -/ +lemma measurable_measure_cond_cdf (ρ : measure (α × ℝ)) : + measurable (λ a, (cond_cdf ρ a).measure) := +begin + rw measure.measurable_measure, + refine λ s hs, measurable_space.induction_on_inter + (borel_eq_generate_from_Iic ℝ) is_pi_system_Iic _ _ _ _ hs, + { simp only [measure_empty, measurable_const], }, + { rintros S ⟨u, rfl⟩, + simp_rw measure_cond_cdf_Iic ρ _ u, + exact (measurable_cond_cdf ρ u).ennreal_of_real, }, + { intros t ht ht_cd_meas, + have : (λ a, (cond_cdf ρ a).measure tᶜ) + = (λ a, (cond_cdf ρ a).measure univ) - (λ a, (cond_cdf ρ a).measure t), + { ext1 a, + rw [measure_compl ht (measure_ne_top (cond_cdf ρ a).measure _), pi.sub_apply], }, + simp_rw [this, measure_cond_cdf_univ ρ], + exact measurable.sub measurable_const ht_cd_meas, }, + { intros f hf_disj hf_meas hf_cd_meas, + simp_rw measure_Union hf_disj hf_meas, + exact measurable.ennreal_tsum hf_cd_meas, }, +end + +end measure + end probability_theory diff --git a/src/probability/kernel/disintegration.lean b/src/probability/kernel/disintegration.lean new file mode 100644 index 0000000000000..d361b61fc6414 --- /dev/null +++ b/src/probability/kernel/disintegration.lean @@ -0,0 +1,411 @@ +/- +Copyright (c) 2023 Rémy Degenne. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Rémy Degenne +-/ +import probability.kernel.cond_cdf +import measure_theory.constructions.polish + +/-! +# Disintegration of measures on product spaces + +Let `ρ` be a finite measure on `α × Ω`, where `Ω` is a standard Borel space. In mathlib terms, `Ω` +verifies `[nonempty Ω] [topological_space Ω] [polish_space Ω] [measurable_space Ω] [borel_space Ω]`. +Then there exists a kernel `ρ.cond_kernel : kernel α Ω` such that for any measurable set +`s : set (α × Ω)`, `ρ s = ∫⁻ a, ρ.cond_kernel a {x | (a, x) ∈ s} ∂ρ.fst`. + +In terms of kernels, `ρ.cond_kernel` is such that for any measurable space `γ`, we +have a disintegration of the constant kernel from `γ` with value `ρ`: +`kernel.const γ ρ = (kernel.const γ ρ.fst) ⊗ₖ (kernel.prod_mk_left γ (cond_kernel ρ))`, +where `ρ.fst` is the marginal measure of `ρ` on `α`. In particular, +`ρ = ((kernel.const unit ρ.fst) ⊗ₖ (kernel.prod_mk_left unit (cond_kernel ρ))) ()`. + +In order to obtain a disintegration for any standard Borel space, we use that these spaces embed +measurably into `ℝ`: it then suffices to define a suitable kernel for `Ω = ℝ`. In the real case, +we define a conditional kernel by taking for each `a : α` the measure associated to the Stieltjes +function `cond_cdf ρ a` (the conditional cumulative distribution function). + +## Main definitions + +* `measure_theory.measure.cond_kernel ρ : kernel α Ω`: conditional kernel described above. + +## Main statements + +* `probability_theory.lintegral_cond_kernel`: + `∫⁻ a, ∫⁻ ω, f (a, ω) ∂(ρ.cond_kernel a) ∂ρ.fst = ∫⁻ x, f x ∂ρ` +* `probability_theory.lintegral_cond_kernel_mem`: + `∫⁻ a, ρ.cond_kernel a {x | (a, x) ∈ s} ∂ρ.fst = ρ s` +* `probability_theory.kernel.const_eq_comp_prod`: + `kernel.const γ ρ = (kernel.const γ ρ.fst) ⊗ₖ (kernel.prod_mk_left γ ρ.cond_kernel)` +* `probability_theory.measure_eq_comp_prod`: + `ρ = ((kernel.const unit ρ.fst) ⊗ₖ (kernel.prod_mk_left unit ρ.cond_kernel)) ()` + +-/ + +open measure_theory set filter + +open_locale ennreal measure_theory topology probability_theory + +namespace probability_theory + +variables {α : Type*} {mα : measurable_space α} + +include mα + +section real + +/-! ### Disintegration of measures on `α × ℝ` -/ + +/-- Conditional measure on the second space of the product given the value on the first, as a +kernel. Use the more general `cond_kernel`. -/ +noncomputable +def cond_kernel_real (ρ : measure (α × ℝ)) : kernel α ℝ := +{ val := λ a, (cond_cdf ρ a).measure, + property := measurable_measure_cond_cdf ρ } + +instance (ρ : measure (α × ℝ)) : is_markov_kernel (cond_kernel_real ρ) := +⟨λ a, by { rw cond_kernel_real, apply_instance, } ⟩ + +lemma cond_kernel_real_Iic (ρ : measure (α × ℝ)) (a : α) (x : ℝ) : + cond_kernel_real ρ a (Iic x) = ennreal.of_real (cond_cdf ρ a x) := +measure_cond_cdf_Iic ρ a x + +lemma set_lintegral_cond_kernel_real_Iic (ρ : measure (α × ℝ)) [is_finite_measure ρ] (x : ℝ) + {s : set α} (hs : measurable_set s) : + ∫⁻ a in s, cond_kernel_real ρ a (Iic x) ∂ρ.fst = ρ (s ×ˢ Iic x) := +by { simp_rw [cond_kernel_real_Iic], exact set_lintegral_cond_cdf ρ x hs, } + +lemma set_lintegral_cond_kernel_real_univ (ρ : measure (α × ℝ)) + {s : set α} (hs : measurable_set s) : + ∫⁻ a in s, cond_kernel_real ρ a univ ∂ρ.fst = ρ (s ×ˢ univ) := +by simp only [measure_univ, lintegral_const, measure.restrict_apply, measurable_set.univ, + univ_inter, one_mul, measure.fst_apply hs, ← prod_univ] + +lemma lintegral_cond_kernel_real_univ (ρ : measure (α × ℝ)) : + ∫⁻ a, cond_kernel_real ρ a univ ∂ρ.fst = ρ univ := +by rw [← set_lintegral_univ, set_lintegral_cond_kernel_real_univ ρ measurable_set.univ, + univ_prod_univ] + +variables (ρ : measure (α × ℝ)) [is_finite_measure ρ] + +lemma set_lintegral_cond_kernel_real_prod + {s : set α} (hs : measurable_set s) {t : set ℝ} (ht : measurable_set t) : + ∫⁻ a in s, cond_kernel_real ρ a t ∂ρ.fst = ρ (s ×ˢ t) := +begin + -- `set_lintegral_cond_kernel_real_Iic` gives the result for `t = Iic x`. These sets form a + -- π-system that generate the borel σ-algebra, hence we can get the same equality for any + -- measurable set `t`. + refine measurable_space.induction_on_inter (borel_eq_generate_from_Iic ℝ) + is_pi_system_Iic _ _ _ _ ht, + { simp only [measure_empty, lintegral_const, zero_mul, prod_empty], }, + { rintros t ⟨q, rfl⟩, + exact set_lintegral_cond_kernel_real_Iic ρ q hs, }, + { intros t ht ht_lintegral, + calc ∫⁻ a in s, cond_kernel_real ρ a tᶜ ∂ρ.fst + = ∫⁻ a in s, (cond_kernel_real ρ a univ) - cond_kernel_real ρ a t ∂ρ.fst : + by { congr' with a, rw measure_compl ht (measure_ne_top (cond_kernel_real ρ a) _), } + ... = ∫⁻ a in s, (cond_kernel_real ρ a univ) ∂ρ.fst - ∫⁻ a in s, cond_kernel_real ρ a t ∂ρ.fst : + begin + rw lintegral_sub (kernel.measurable_coe (cond_kernel_real ρ) ht), + { rw ht_lintegral, + exact measure_ne_top ρ _, }, + { exact eventually_of_forall (λ a, measure_mono (subset_univ _)), }, + end + ... = ρ (s ×ˢ univ) - ρ (s ×ˢ t) : + by rw [set_lintegral_cond_kernel_real_univ ρ hs, ht_lintegral] + ... = ρ (s ×ˢ tᶜ) : + begin + rw ← measure_diff _ (hs.prod ht) (measure_ne_top ρ _), + { rw [prod_diff_prod, compl_eq_univ_diff], + simp only [diff_self, empty_prod, union_empty], }, + { rw prod_subset_prod_iff, + exact or.inl ⟨subset_rfl, subset_univ t⟩, }, + end, }, + { intros f hf_disj hf_meas hf_eq, + simp_rw measure_Union hf_disj hf_meas, + rw [lintegral_tsum (λ i, (kernel.measurable_coe _ (hf_meas i)).ae_measurable.restrict), + prod_Union, measure_Union], + { simp_rw hf_eq, }, + { intros i j hij, + rw [function.on_fun, disjoint_prod], + exact or.inr (hf_disj hij), }, + { exact λ i, measurable_set.prod hs (hf_meas i), }, }, +end + +lemma lintegral_cond_kernel_real_mem {s : set (α × ℝ)} (hs : measurable_set s) : + ∫⁻ a, cond_kernel_real ρ a {x | (a, x) ∈ s} ∂ρ.fst = ρ s := +begin + -- `set_lintegral_cond_kernel_real_prod` gives the result for sets of the form `t₁ × t₂`. These + -- sets form a π-system that generate the product σ-algebra, hence we can get the same equality + -- for any measurable set `s`. + refine measurable_space.induction_on_inter generate_from_prod.symm is_pi_system_prod _ _ _ _ hs, + { simp only [mem_empty_iff_false, set_of_false, measure_empty, lintegral_const, zero_mul], }, + { intros t ht, + rw mem_image2 at ht, + obtain ⟨t₁, t₂, ht₁, ht₂, rfl⟩ := ht, + have h_prod_eq_snd : ∀ a ∈ t₁, {x : ℝ | (a, x) ∈ t₁ ×ˢ t₂} = t₂, + { intros a ha, + simp only [ha, prod_mk_mem_set_prod_eq, true_and, set_of_mem_eq], }, + cases eq_empty_or_nonempty t₂ with h h, + { simp only [h, prod_empty, mem_empty_iff_false, set_of_false, measure_empty, lintegral_const, + zero_mul], }, + rw ← lintegral_add_compl _ ht₁, + have h_eq1 : ∫⁻ a in t₁, cond_kernel_real ρ a {x : ℝ | (a, x) ∈ t₁ ×ˢ t₂} ∂ρ.fst + = ∫⁻ a in t₁, cond_kernel_real ρ a t₂ ∂ρ.fst, + { refine set_lintegral_congr_fun ht₁ (eventually_of_forall (λ a ha, _)), + rw h_prod_eq_snd a ha, }, + have h_eq2 : ∫⁻ a in t₁ᶜ, cond_kernel_real ρ a {x : ℝ | (a, x) ∈ t₁ ×ˢ t₂} ∂ρ.fst = 0, + { suffices h_eq_zero : ∀ a ∈ t₁ᶜ, cond_kernel_real ρ a {x : ℝ | (a, x) ∈ t₁ ×ˢ t₂} = 0, + { rw set_lintegral_congr_fun ht₁.compl (eventually_of_forall h_eq_zero), + simp only [lintegral_const, zero_mul], }, + intros a hat₁, + rw mem_compl_iff at hat₁, + simp only [hat₁, prod_mk_mem_set_prod_eq, false_and, set_of_false, measure_empty], }, + rw [h_eq1, h_eq2, add_zero], + exact set_lintegral_cond_kernel_real_prod ρ ht₁ ht₂, }, + { intros t ht ht_eq, + calc ∫⁻ a, cond_kernel_real ρ a {x : ℝ | (a, x) ∈ tᶜ} ∂ρ.fst + = ∫⁻ a, cond_kernel_real ρ a {x : ℝ | (a, x) ∈ t}ᶜ ∂ρ.fst : rfl + ... = ∫⁻ a, cond_kernel_real ρ a univ - cond_kernel_real ρ a {x : ℝ | (a, x) ∈ t} ∂ρ.fst : + begin + congr' with a : 1, + exact measure_compl (measurable_prod_mk_left ht) (measure_ne_top (cond_kernel_real ρ a) _), + end + ... = ∫⁻ a, cond_kernel_real ρ a univ ∂ρ.fst + - ∫⁻ a, cond_kernel_real ρ a {x : ℝ | (a, x) ∈ t} ∂ρ.fst : + begin + have h_le : (λ a, cond_kernel_real ρ a {x : ℝ | (a, x) ∈ t}) + ≤ᵐ[ρ.fst] λ a, cond_kernel_real ρ a univ, + { exact eventually_of_forall (λ a, measure_mono (subset_univ _)), }, + rw lintegral_sub _ _ h_le, + { exact kernel.measurable_kernel_prod_mk_left ht, }, + refine ((lintegral_mono_ae h_le).trans_lt _).ne, + rw lintegral_cond_kernel_real_univ, + exact measure_lt_top ρ univ, + end + ... = ρ univ - ρ t : by rw [ht_eq, lintegral_cond_kernel_real_univ] + ... = ρ tᶜ : (measure_compl ht (measure_ne_top _ _)).symm, }, + { intros f hf_disj hf_meas hf_eq, + have h_eq : ∀ a, {x | (a, x) ∈ ⋃ i, f i} = ⋃ i, {x | (a, x) ∈ f i}, + { intros a, + ext1 x, + simp only [mem_Union, mem_set_of_eq], }, + simp_rw h_eq, + have h_disj : ∀ a, pairwise (disjoint on (λ i, {x | (a, x) ∈ f i})), + { intros a i j hij, + have h_disj := hf_disj hij, + rw [function.on_fun, disjoint_iff_inter_eq_empty] at h_disj ⊢, + ext1 x, + simp only [mem_inter_iff, mem_set_of_eq, mem_empty_iff_false, iff_false], + intros h_mem_both, + suffices : (a, x) ∈ ∅, by rwa mem_empty_iff_false at this, + rwa [← h_disj, mem_inter_iff], }, + calc ∫⁻ a, cond_kernel_real ρ a (⋃ i, {x | (a, x) ∈ f i}) ∂ρ.fst + = ∫⁻ a, ∑' i, cond_kernel_real ρ a {x | (a, x) ∈ f i} ∂ρ.fst : + by { congr' with a : 1, + rw measure_Union (h_disj a) (λ i, measurable_prod_mk_left (hf_meas i)), } + ... = ∑' i, ∫⁻ a, cond_kernel_real ρ a {x | (a, x) ∈ f i} ∂ρ.fst : lintegral_tsum + (λ i, (kernel.measurable_kernel_prod_mk_left (hf_meas i)).ae_measurable) + ... = ∑' i, ρ (f i) : by simp_rw hf_eq + ... = ρ (Union f) : (measure_Union hf_disj hf_meas).symm, }, +end + +theorem kernel.const_eq_comp_prod_real (γ : Type*) [measurable_space γ] + (ρ : measure (α × ℝ)) [is_finite_measure ρ] : + kernel.const γ ρ = (kernel.const γ ρ.fst) ⊗ₖ (kernel.prod_mk_left γ (cond_kernel_real ρ)) := +begin + ext a s hs : 2, + rw [kernel.comp_prod_apply _ _ _ hs, kernel.const_apply, kernel.const_apply], + simp_rw kernel.prod_mk_left_apply, + rw lintegral_cond_kernel_real_mem ρ hs, +end + +theorem measure_eq_comp_prod_real : + ρ = ((kernel.const unit ρ.fst) ⊗ₖ (kernel.prod_mk_left unit (cond_kernel_real ρ))) () := +by rw [← kernel.const_eq_comp_prod_real unit ρ, kernel.const_apply] + +lemma lintegral_cond_kernel_real {f : α × ℝ → ℝ≥0∞} (hf : measurable f) : + ∫⁻ a, ∫⁻ y, f (a, y) ∂(cond_kernel_real ρ a) ∂ρ.fst = ∫⁻ x, f x ∂ρ := +begin + nth_rewrite 1 measure_eq_comp_prod_real ρ, + rw [kernel.lintegral_comp_prod _ _ _ hf, kernel.const_apply], + simp_rw kernel.prod_mk_left_apply, +end + +lemma ae_cond_kernel_real_eq_one {s : set ℝ} (hs : measurable_set s) (hρ : ρ {x | x.snd ∈ sᶜ} = 0) : + ∀ᵐ a ∂ρ.fst, cond_kernel_real ρ a s = 1 := +begin + have h : ρ {x | x.snd ∈ sᶜ} + = (kernel.const unit ρ.fst ⊗ₖ kernel.prod_mk_left unit (cond_kernel_real ρ)) () + {x | x.snd ∈ sᶜ}, + { rw ← measure_eq_comp_prod_real, }, + rw [hρ, kernel.comp_prod_apply] at h, + swap, { exact measurable_snd hs.compl, }, + rw [eq_comm, lintegral_eq_zero_iff] at h, + swap, + { simp_rw kernel.prod_mk_left_apply', + simp only [mem_compl_iff, mem_set_of_eq], + exact kernel.measurable_coe _ hs.compl, }, + rw kernel.const_apply at h, + simp only [mem_compl_iff, mem_set_of_eq, kernel.prod_mk_left_apply'] at h, + filter_upwards [h] with a ha, + change cond_kernel_real ρ a sᶜ = 0 at ha, + rwa [prob_compl_eq_zero_iff hs] at ha, + apply_instance, +end + +end real + +section polish + +/-! ### Disintegration of measures on Polish Borel spaces + +Since every standard Borel space embeds measurably into `ℝ`, we can generalize the disintegration +property on `ℝ` to all these spaces. -/ + +variables {Ω : Type*} [topological_space Ω] [polish_space Ω] [measurable_space Ω] [borel_space Ω] + [nonempty Ω] + (ρ : measure (α × Ω)) [is_finite_measure ρ] + +/-- Existence of a conditional kernel. Use the definition `cond_kernel` to get that kernel. -/ +lemma exists_cond_kernel (γ : Type*) [measurable_space γ] : + ∃ (η : kernel α Ω) (h : is_markov_kernel η), + kernel.const γ ρ + = @kernel.comp_prod γ α _ _ Ω _ (kernel.const γ ρ.fst) _ (kernel.prod_mk_left γ η) + (by { haveI := h, apply_instance, }) := +begin + obtain ⟨f, hf⟩ := exists_measurable_embedding_real Ω, + let ρ' : measure (α × ℝ) := ρ.map (prod.map id f), + -- The general idea is to define `η = kernel.comap_right (cond_kernel_real ρ') hf`. There is + -- however an issue: that `η` may not be a Markov kernel since its value is only a + -- probability distribution almost everywhere with respect to `ρ.fst`, not everywhere. + -- We modify it to obtain a Markov kernel which is almost everywhere equal. + let ρ_set := (to_measurable ρ.fst {a | cond_kernel_real ρ' a (range f) = 1}ᶜ)ᶜ, + have hm : measurable_set ρ_set := (measurable_set_to_measurable _ _).compl, + have h_eq_one_of_mem : ∀ a ∈ ρ_set, cond_kernel_real ρ' a (range f) = 1, + { intros a ha, + rw [mem_compl_iff] at ha, + have h_ss := subset_to_measurable ρ.fst {a : α | cond_kernel_real ρ' a (range f) = 1}ᶜ, + suffices ha' : a ∉ {a : α | cond_kernel_real ρ' a (range f) = 1}ᶜ, + { rwa not_mem_compl_iff at ha', }, + exact not_mem_subset h_ss ha, }, + have h_prod_embed : measurable_embedding (prod.map (id : α → α) f) := + (measurable_embedding.id).prod_mk hf, + have h_fst : ρ'.fst = ρ.fst, + { ext1 u hu, + rw [measure.fst_apply hu, measure.fst_apply hu, + measure.map_apply h_prod_embed.measurable (measurable_fst hu)], + refl, }, + have h_ae : ∀ᵐ a ∂ρ.fst, a ∈ ρ_set, + { rw ae_iff, + simp only [not_mem_compl_iff, set_of_mem_eq, measure_to_measurable], + change (ρ.fst) {a : α | a ∉ {a' : α | cond_kernel_real ρ' a' (range f) = 1}} = 0, + rw [← ae_iff, ← h_fst], + refine ae_cond_kernel_real_eq_one ρ' hf.measurable_set_range _, + rw measure.map_apply h_prod_embed.measurable, + swap, { exact measurable_snd hf.measurable_set_range.compl, }, + convert measure_empty, + ext1 x, + simp only [mem_compl_iff, mem_range, preimage_set_of_eq, prod_map, mem_set_of_eq, + mem_empty_iff_false, iff_false, not_not, exists_apply_eq_apply], }, + classical, + obtain ⟨x₀, hx₀⟩ : ∃ x, x ∈ range f := range_nonempty _, + let η' := kernel.piecewise hm (cond_kernel_real ρ') + (kernel.deterministic (λ _, x₀) measurable_const), + -- We show that `kernel.comap_right η' hf` is a suitable Markov kernel. + refine ⟨kernel.comap_right η' hf, _, _⟩, + { refine kernel.is_markov_kernel.comap_right _ _ (λ a, _), + rw kernel.piecewise_apply', + split_ifs with h_mem h_not_mem, + { exact h_eq_one_of_mem _ h_mem, }, + { rw [kernel.deterministic_apply' _ _ hf.measurable_set_range, set.indicator_apply, + if_pos hx₀], }, }, + have : kernel.const γ ρ = kernel.comap_right (kernel.const γ ρ') h_prod_embed, + { ext c t ht : 2, + rw [kernel.const_apply, kernel.comap_right_apply' _ _ _ ht, kernel.const_apply, + measure.map_apply h_prod_embed.measurable (h_prod_embed.measurable_set_image.mpr ht)], + congr' with x : 1, + rw ← @prod.mk.eta _ _ x, + simp only [id.def, mem_preimage, prod.map_mk, mem_image, prod.mk.inj_iff, prod.exists], + refine ⟨λ h, ⟨x.1, x.2, h, rfl, rfl⟩, _⟩, + rintros ⟨a, b, h_mem, rfl, hf_eq⟩, + rwa hf.injective hf_eq at h_mem, }, + rw [this, kernel.const_eq_comp_prod_real _ ρ'], + ext c t ht : 2, + rw [kernel.comap_right_apply' _ _ _ ht, + kernel.comp_prod_apply _ _ _ (h_prod_embed.measurable_set_image.mpr ht), kernel.const_apply, + h_fst, kernel.comp_prod_apply _ _ _ ht, kernel.const_apply], + refine lintegral_congr_ae _, + filter_upwards [h_ae] with a ha, + rw [kernel.prod_mk_left_apply', kernel.prod_mk_left_apply', kernel.comap_right_apply'], + swap, { exact measurable_prod_mk_left ht, }, + have h1 : {c : ℝ | (a, c) ∈ prod.map id f '' t} = f '' {c : Ω | (a, c) ∈ t}, + { ext1 x, + simp only [prod_map, id.def, mem_image, prod.mk.inj_iff, prod.exists, mem_set_of_eq], + split, + { rintros ⟨a', b, h_mem, rfl, hf_eq⟩, + exact ⟨b, h_mem, hf_eq⟩, }, + { rintros ⟨b, h_mem, hf_eq⟩, + exact ⟨a, b, h_mem, rfl, hf_eq⟩, }, }, + have h2 : cond_kernel_real ρ' (c, a).snd = η' (c, a).snd, + { rw [kernel.piecewise_apply, if_pos ha], }, + rw [h1, h2], +end + +/-- Conditional kernel of a measure on a product space: a Markov kernel such that +`ρ = ((kernel.const unit ρ.fst) ⊗ₖ (kernel.prod_mk_left unit ρ.cond_kernel)) ()` +(see `probability_theory.measure_eq_comp_prod`). -/ +@[irreducible] noncomputable +def _root_.measure_theory.measure.cond_kernel : kernel α Ω := +(exists_cond_kernel ρ unit).some + +lemma cond_kernel_def : ρ.cond_kernel = (exists_cond_kernel ρ unit).some := +by rw measure_theory.measure.cond_kernel + +instance : is_markov_kernel ρ.cond_kernel := +by { rw cond_kernel_def, exact (exists_cond_kernel ρ unit).some_spec.some, } + +lemma kernel.const_unit_eq_comp_prod : + kernel.const unit ρ + = (kernel.const unit ρ.fst) ⊗ₖ (kernel.prod_mk_left unit ρ.cond_kernel) := +by { simp_rw cond_kernel_def, exact (exists_cond_kernel ρ unit).some_spec.some_spec, } + +/-- **Disintegration** of finite product measures on `α × Ω`, where `Ω` is Polish Borel. Such a +measure can be written as the composition-product of the constant kernel with value `ρ.fst` +(marginal measure over `α`) and a Markov kernel from `α` to `Ω`. We call that Markov kernel +`probability_theory.cond_kernel ρ`. -/ +theorem measure_eq_comp_prod : + ρ = ((kernel.const unit ρ.fst) ⊗ₖ (kernel.prod_mk_left unit ρ.cond_kernel)) () := +by rw [← kernel.const_unit_eq_comp_prod, kernel.const_apply] + +/-- **Disintegration** of constant kernels. A constant kernel on a product space `α × Ω`, where `Ω` +is Polish Borel, can be written as the composition-product of the constant kernel with value `ρ.fst` +(marginal measure over `α`) and a Markov kernel from `α` to `Ω`. We call that Markov kernel +`probability_theory.cond_kernel ρ`. -/ +theorem kernel.const_eq_comp_prod (γ : Type*) [measurable_space γ] + (ρ : measure (α × Ω)) [is_finite_measure ρ] : + kernel.const γ ρ = (kernel.const γ ρ.fst) ⊗ₖ (kernel.prod_mk_left γ ρ.cond_kernel) := +begin + ext a s hs : 2, + simpa only [kernel.const_apply, kernel.comp_prod_apply _ _ _ hs, kernel.prod_mk_left_apply'] + using kernel.ext_iff'.mp (kernel.const_unit_eq_comp_prod ρ) () s hs, +end + +lemma lintegral_cond_kernel_mem {s : set (α × Ω)} (hs : measurable_set s) : + ∫⁻ a, ρ.cond_kernel a {x | (a, x) ∈ s} ∂ρ.fst = ρ s := +begin + conv_rhs { rw measure_eq_comp_prod ρ, }, + simp_rw [kernel.comp_prod_apply _ _ _ hs, kernel.const_apply, kernel.prod_mk_left_apply], +end + +lemma lintegral_cond_kernel {f : α × Ω → ℝ≥0∞} (hf : measurable f) : + ∫⁻ a, ∫⁻ ω, f (a, ω) ∂(ρ.cond_kernel a) ∂ρ.fst = ∫⁻ x, f x ∂ρ := +begin + conv_rhs { rw measure_eq_comp_prod ρ, }, + rw [kernel.lintegral_comp_prod _ _ _ hf, kernel.const_apply], + simp_rw kernel.prod_mk_left_apply, +end + +end polish + +end probability_theory From b353176c24d96c23f0ce1cc63efc3f55019702d9 Mon Sep 17 00:00:00 2001 From: Eric Rodriguez Date: Mon, 22 May 2023 12:50:02 +0000 Subject: [PATCH 1053/1291] feat(field_theory/splitting_field): fix the diamond (#18857) This re-implements how instances for `splitting_field` in such a way that we can now avoid diamonds with the N, Z and Q-actions. This makes a lot more instances safe, which is good for e.g. flt-regular. Many thanks to @Vierkantor for many of the ideas and all the `distrib_smul` PRs that were needed for this. Co-authored-by: Eric Wieser Co-authored-by: Riccardo Brasca --- src/field_theory/splitting_field.lean | 305 +++++++++++++++++++----- src/number_theory/cyclotomic/basic.lean | 9 +- src/number_theory/cyclotomic/rat.lean | 20 +- 3 files changed, 254 insertions(+), 80 deletions(-) diff --git a/src/field_theory/splitting_field.lean b/src/field_theory/splitting_field.lean index 50eb8508117e3..cf3ef1028d045 100644 --- a/src/field_theory/splitting_field.lean +++ b/src/field_theory/splitting_field.lean @@ -3,6 +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 -/ +import algebra.char_p.algebra import field_theory.intermediate_field import ring_theory.adjoin.field @@ -109,46 +110,234 @@ namespace splitting_field_aux theorem succ (n : ℕ) (f : K[X]) : splitting_field_aux (n+1) f = splitting_field_aux n f.remove_factor := rfl -instance field (n : ℕ) : Π {K : Type u} [field K], by exactI - Π {f : K[X]}, field (splitting_field_aux n f) := -nat.rec_on n (λ K _ _, ‹field K›) $ λ n ih K _ f, ih +section lift_instances -instance inhabited {n : ℕ} {f : K[X]} : - inhabited (splitting_field_aux n f) := ⟨37⟩ +/-! ### Instances on `splitting_field_aux` -/- -Note that the recursive nature of this definition and `splitting_field_aux.field` creates -non-definitionally-equal diamonds in the `ℕ`- and `ℤ`- actions. -```lean -example (n : ℕ) {K : Type u} [field K] {f : K[X]} (hfn : f.nat_degree = n) : - (add_comm_monoid.nat_module : module ℕ (splitting_field_aux n f hfn)) = - @algebra.to_module _ _ _ _ (splitting_field_aux.algebra n _ hfn) := -rfl -- fails -``` -It's not immediately clear whether this _can_ be fixed; the failure is much the same as the reason -that the following fails: -```lean -def cases_twice {α} (a₀ aₙ : α) : ℕ → α × α -| 0 := (a₀, a₀) -| (n + 1) := (aₙ, aₙ) - -example (x : ℕ) {α} (a₀ aₙ : α) : (cases_twice a₀ aₙ x).1 = (cases_twice a₀ aₙ x).2 := rfl -- fails -``` -We don't really care at this point because this is an implementation detail (which is why this is -not a docstring), but we do in `splitting_field.algebra'` below. -/ -instance algebra (n : ℕ) : Π (R : Type*) {K : Type u} [comm_semiring R] [field K], - by exactI Π [algebra R K] {f : K[X]}, - algebra R (splitting_field_aux n f) := -nat.rec_on n (λ R K _ _ _ _, by exactI ‹algebra R K›) $ - λ n ih R K _ _ _ f, by exactI ih R +In order to avoid diamond issues, we have to be careful to define each data field of algebraic +instances on `splitting_field_aux` by recursion, rather than defining the whole structure by +recursion at once. + +The important detail is that the `smul` instances can be lifted _before_ we create the algebraic +instances; if we do not do this, this creates a mutual dependence on both on each other, and it +is impossible to untangle this mess. In order to do this, we need that these `smul` instances +are distributive in the sense of `distrib_smul`, which is weak enough to be guaranteed at this +stage. This is sufficient to lift an action to `adjoin_root f` (remember that this is a quotient, +so these conditions are equivalent to well-definedness), which is all we need. +-/ + +/-- Splitting fields have a zero. -/ +protected def zero (n : ℕ) : Π {K : Type u} [field K], by exactI Π {f : K[X]}, + splitting_field_aux n f := +nat.rec_on n (λ K fK f, by exactI @has_zero.zero K _) (λ n ih K fK f, ih) + +/-- Splitting fields have an addition. -/ +protected def add (n : ℕ) : Π {K : Type u} [field K], by exactI Π {f : K[X]}, + splitting_field_aux n f → splitting_field_aux n f → splitting_field_aux n f := +nat.rec_on n (λ K fK f, by exactI @has_add.add K _) (λ n ih K fK f, ih) + +/-- Splitting fields inherit scalar multiplication. -/ +protected def smul (n : ℕ) : Π (α : Type*) {K : Type u} [field K], + by exactI Π [distrib_smul α K], + by exactI Π [is_scalar_tower α K K] {f : K[X]}, + α → splitting_field_aux n f → splitting_field_aux n f := +nat.rec_on n + (λ α K fK ds ist f, by exactI @has_smul.smul _ K _) + (λ n ih α K fK ds ist f, by exactI ih α) + +instance has_smul (α : Type*) (n : ℕ) {K : Type u} [field K] [distrib_smul α K] + [is_scalar_tower α K K] {f : K[X]} : + has_smul α (splitting_field_aux n f) := +⟨splitting_field_aux.smul n α⟩ instance is_scalar_tower (n : ℕ) : Π (R₁ R₂ : Type*) {K : Type u} - [comm_semiring R₁] [comm_semiring R₂] [has_smul R₁ R₂] [field K], - by exactI Π [algebra R₁ K] [algebra R₂ K], - by exactI Π [is_scalar_tower R₁ R₂ K] {f : K[X]}, - is_scalar_tower R₁ R₂ (splitting_field_aux n f) := -nat.rec_on n (λ R₁ R₂ K _ _ _ _ _ _ _ _, by exactI ‹is_scalar_tower R₁ R₂ K›) $ - λ n ih R₁ R₂ K _ _ _ _ _ _ _ f, by exactI ih R₁ R₂ + [has_smul R₁ R₂] [field K], + by exactI Π [distrib_smul R₂ K] [distrib_smul R₁ K], + by exactI Π [is_scalar_tower R₁ K K] [is_scalar_tower R₂ K K] [is_scalar_tower R₁ R₂ K] + {f : K[X]}, by exactI is_scalar_tower R₁ R₂ (splitting_field_aux n f) := +nat.rec_on n (λ R₁ R₂ K _ _ hs₂ hs₁ _ _ h f, +begin + rcases hs₁ with @⟨@⟨⟨hs₁⟩,_⟩,_⟩, + rcases hs₂ with @⟨@⟨⟨hs₂⟩,_⟩,_⟩, + exact h, +end) $ λ n ih R₁ R₂ K _ _ _ _ _ _ _ f, by exactI ih R₁ R₂ + +/-- Splitting fields have a negation. -/ +protected def neg (n : ℕ) : Π {K : Type u} [field K], by exactI Π {f : K[X]}, + splitting_field_aux n f → splitting_field_aux n f := +nat.rec_on n (λ K fK f, by exactI @has_neg.neg K _) (λ n ih K fK f, ih) + +/-- Splitting fields have subtraction. -/ +protected def sub (n : ℕ) : Π {K : Type u} [field K], by exactI Π {f : K[X]}, + splitting_field_aux n f → splitting_field_aux n f → splitting_field_aux n f := +nat.rec_on n (λ K fK f, by exactI @has_sub.sub K _) (λ n ih K fK f, ih) + +/-- Splitting fields have a one. -/ +protected def one (n : ℕ) : Π {K : Type u} [field K], by exactI Π {f : K[X]}, + splitting_field_aux n f := +nat.rec_on n (λ K fK f, by exactI @has_one.one K _) (λ n ih K fK f, ih) + +/-- Splitting fields have a multiplication. -/ +protected def mul (n : ℕ) : Π {K : Type u} [field K], by exactI Π {f : K[X]}, + splitting_field_aux n f → splitting_field_aux n f → splitting_field_aux n f := +nat.rec_on n (λ K fK f, by exactI @has_mul.mul K _) (λ n ih K fK f, ih) + +/-- Splitting fields have a power operation. -/ +protected def npow (n : ℕ) : Π {K : Type u} [field K], by exactI Π {f : K[X]}, + ℕ → splitting_field_aux n f → splitting_field_aux n f := +nat.rec_on n (λ K fK f n x, by exactI @has_pow.pow K _ _ x n) (λ n ih K fK f, ih) + +/-- Splitting fields have an injection from the base field. -/ +protected def mk (n : ℕ) : Π {K : Type u} [field K], + by exactI Π {f : K[X]}, K → splitting_field_aux n f := +nat.rec_on n (λ K fK f, id) (λ n ih K fK f, by exactI ih ∘ coe) +-- note that `coe` goes from `K → adjoin_root f`, and then `ih` lifts to the full splitting field +-- (as we generalise over all fields in this recursion!) + +/-- Splitting fields have an inverse. -/ +protected def inv (n : ℕ) : Π {K : Type u} [field K], by exactI Π {f : K[X]}, + splitting_field_aux n f → splitting_field_aux n f := +nat.rec_on n (λ K fK f, by exactI @has_inv.inv K _) (λ n ih K fK f, ih) + +/-- Splitting fields have a division. -/ +protected def div (n : ℕ) : Π {K : Type u} [field K], by exactI Π {f : K[X]}, + splitting_field_aux n f → splitting_field_aux n f → splitting_field_aux n f := +nat.rec_on n (λ K fK f, by exactI @has_div.div K _) (λ n ih K fK f, ih) + +/-- Splitting fields have powers by integers. -/ +protected def zpow (n : ℕ) : Π {K : Type u} [field K], by exactI Π {f : K[X]}, + ℤ → splitting_field_aux n f → splitting_field_aux n f := +nat.rec_on n (λ K fK f n x, by exactI @has_pow.pow K _ _ x n) (λ n ih K fK f, ih) + +-- I'm not sure why these two lemmas break, but inlining them seems to not work. +private lemma nsmul_zero (n : ℕ) : Π {K : Type u} [field K], by exactI Π {f : K[X]} + (x : splitting_field_aux n f), (0 : ℕ) • x = splitting_field_aux.zero n := +nat.rec_on n (λ K fK f, by exactI zero_nsmul) (λ n ih K fK f, by exactI ih) + +private lemma nsmul_succ (n : ℕ) : Π {K : Type u} [field K], by exactI Π {f : K[X]} + (k : ℕ) (x : splitting_field_aux n f), + (k + 1) • x = splitting_field_aux.add n x (k • x) := +nat.rec_on n (λ K fK f n x, by exactI succ_nsmul x n) (λ n ih K fK f, by exactI ih) + +instance field (n : ℕ) {K : Type u} [field K] {f : K[X]} : + field (splitting_field_aux n f) := +begin + refine + { zero := splitting_field_aux.zero n, + one := splitting_field_aux.one n, + add := splitting_field_aux.add n, + zero_add := have h : _ := @zero_add, _, + add_zero := have h : _ := @add_zero, _, + add_assoc := have h : _ := @add_assoc, _, + add_comm := have h : _ := @add_comm, _, + neg := splitting_field_aux.neg n, + add_left_neg := have h : _ := @add_left_neg, _, + sub := splitting_field_aux.sub n, + sub_eq_add_neg := have h : _ := @sub_eq_add_neg, _, + mul := splitting_field_aux.mul n, + one_mul := have h : _ := @one_mul, _, + mul_one := have h : _ := @mul_one, _, + mul_assoc := have h : _ := @mul_assoc, _, + left_distrib := have h : _ := @left_distrib, _, + right_distrib := have h : _ := @right_distrib, _, + mul_comm := have h : _ := @mul_comm, _, + inv := splitting_field_aux.inv n, + inv_zero := have h : _ := @inv_zero, _, + div := splitting_field_aux.div n, + div_eq_mul_inv := have h : _ := @div_eq_mul_inv, _, + mul_inv_cancel := have h : _ := @mul_inv_cancel, _, + exists_pair_ne := have h : _ := @exists_pair_ne, _, + nat_cast := splitting_field_aux.mk n ∘ (coe : ℕ → K), + nat_cast_zero := have h : _ := @comm_ring.nat_cast_zero, _, + nat_cast_succ := have h : _ := @comm_ring.nat_cast_succ, _, + nsmul := (•), + nsmul_zero' := nsmul_zero n, + nsmul_succ' := nsmul_succ n, + int_cast := splitting_field_aux.mk n ∘ (coe : ℤ → K), + int_cast_of_nat := have h : _ := @comm_ring.int_cast_of_nat, _, + int_cast_neg_succ_of_nat := have h : _ := @comm_ring.int_cast_neg_succ_of_nat, _, + zsmul := (•), + zsmul_zero' := have h : _ := @subtraction_comm_monoid.zsmul_zero', _, + zsmul_succ' := have h : _ := @subtraction_comm_monoid.zsmul_succ', _, + zsmul_neg' := have h : _ := @subtraction_comm_monoid.zsmul_neg', _, + rat_cast := splitting_field_aux.mk n ∘ (coe : ℚ → K), + rat_cast_mk := have h : _ := @field.rat_cast_mk, _, + qsmul := (•), + qsmul_eq_mul' := have h : _ := @field.qsmul_eq_mul', _, + npow := splitting_field_aux.npow n, + npow_zero' := have h : _ := @comm_ring.npow_zero', _, + npow_succ' := have h : _ := @comm_ring.npow_succ', _, + zpow := splitting_field_aux.zpow n, + zpow_zero' := have h : _ := @field.zpow_zero', _, + zpow_succ' := have h : _ := @field.zpow_succ', _, + zpow_neg' := have h : _ := @field.zpow_neg', _}, + all_goals { + unfreezingI { induction n with n ih generalizing K }, + { apply @h K }, + { exact @ih _ _ _ } }, +end + +instance inhabited {n : ℕ} {f : K[X]} : + inhabited (splitting_field_aux n f) := ⟨37⟩ + +/-- The injection from the base field as a ring homomorphism. -/ +@[simps] protected def mk_hom (n : ℕ) {K : Type u} [field K] {f : K[X]} : + K →+* splitting_field_aux n f := +{ to_fun := splitting_field_aux.mk n, + map_one' := + begin + unfreezingI { induction n with k hk generalizing K }, + { simp [splitting_field_aux.mk] }, + exact hk + end, + map_mul' := + begin + unfreezingI { induction n with k hk generalizing K }, + { simp [splitting_field_aux.mk] }, + intros x y, + change (splitting_field_aux.mk k) ((adjoin_root.of f.factor) _) = _, + rw [map_mul], + exact hk _ _ + end, + map_zero' := + begin + unfreezingI { induction n with k hk generalizing K }, + { simp [splitting_field_aux.mk] }, + change (splitting_field_aux.mk k) ((adjoin_root.of f.factor) 0) = _, + rw [map_zero, hk], + end, + map_add' := + begin + unfreezingI { induction n with k hk generalizing K }, + { simp [splitting_field_aux.mk] }, + intros x y, + change (splitting_field_aux.mk k) ((adjoin_root.of f.factor) _) = _, + rw [map_add], + exact hk _ _ + end } + +instance algebra (n : ℕ) (R : Type*) {K : Type u} [comm_semiring R] [field K] + [algebra R K] {f : K[X]} : algebra R (splitting_field_aux n f) := +{ to_fun := (splitting_field_aux.mk_hom n).comp (algebra_map R K), + smul := (•), + smul_def' := + begin + unfreezingI { induction n with k hk generalizing K }, + { exact algebra.smul_def }, + exact hk + end, + commutes' := λ a b, mul_comm _ _, + .. (splitting_field_aux.mk_hom n).comp (algebra_map R K) } + +/-- Because `splitting_field_aux` is defined by recursion, we have to make sure all instances +on `splitting_field_aux` are defined by recursion within the fields. Otherwise, there will be +instance diamonds such as the following: -/ +example (n : ℕ) {K : Type u} [field K] {f : K[X]} : + (add_comm_monoid.nat_module : module ℕ (splitting_field_aux n f)) = + @algebra.to_module _ _ _ _ (splitting_field_aux.algebra n _) := +rfl + +end lift_instances instance algebra''' {n : ℕ} {f : K[X]} : algebra (adjoin_root f.factor) @@ -248,36 +437,32 @@ splitting_field_aux.field _ instance inhabited : inhabited (splitting_field f) := ⟨37⟩ -/-- This should be an instance globally, but it creates diamonds with the `ℕ`, `ℤ`, and `ℚ` algebras -(via their `smul` and `to_fun` fields): +instance algebra' {R} [comm_semiring R] [algebra R K] : algebra R (splitting_field f) := +splitting_field_aux.algebra _ _ -```lean -example : - (algebra_nat : algebra ℕ (splitting_field f)) = splitting_field.algebra' f := -rfl -- fails +instance : algebra K (splitting_field f) := +splitting_field_aux.algebra _ _ -example : - (algebra_int _ : algebra ℤ (splitting_field f)) = splitting_field.algebra' f := -rfl -- fails +instance [char_zero K] : char_zero (splitting_field f) := +char_zero_of_injective_algebra_map ((algebra_map K _).injective) -example [char_zero K] [char_zero (splitting_field f)] : - (algebra_rat : algebra ℚ (splitting_field f)) = splitting_field.algebra' f := -rfl -- fails -``` +-- The algebra instance deriving from `K` should be definitionally equal to that +-- deriving from the field structure on `splitting_field f`. +example : (add_comm_monoid.nat_module : module ℕ (splitting_field f)) = + @algebra.to_module _ _ _ _ (splitting_field.algebra' f) := +rfl -Until we resolve these diamonds, it's more convenient to only turn this instance on with -`local attribute [instance]` in places where the benefit of having the instance outweighs the cost. +example : (add_comm_group.int_module _ : module ℤ (splitting_field f)) = + @algebra.to_module _ _ _ _ (splitting_field.algebra' f) := +rfl -In the meantime, the `splitting_field.algebra` instance below is immune to these particular diamonds -since `K = ℕ` and `K = ℤ` are not possible due to the `field K` assumption. Diamonds in -`algebra ℚ (splitting_field f)` instances are still possible via this instance unfortunately, but -these are less common as they require suitable `char_zero` instances to be present. --/ -instance algebra' {R} [comm_semiring R] [algebra R K] : algebra R (splitting_field f) := -splitting_field_aux.algebra _ _ +example [char_zero K] : (splitting_field.algebra' f) = algebra_rat := +rfl -instance : algebra K (splitting_field f) := -splitting_field_aux.algebra _ _ +example [char_zero K] : (splitting_field.algebra' f) = algebra_rat := +rfl + +example {q : ℚ[X]} : algebra_int (splitting_field q) = splitting_field.algebra' q := rfl protected theorem splits : splits (algebra_map K (splitting_field f)) f := splitting_field_aux.splits _ _ rfl diff --git a/src/number_theory/cyclotomic/basic.lean b/src/number_theory/cyclotomic/basic.lean index 05c071519e7f2..56086e1890cfe 100644 --- a/src/number_theory/cyclotomic/basic.lean +++ b/src/number_theory/cyclotomic/basic.lean @@ -553,12 +553,13 @@ variables [is_domain A] [algebra A K] [is_fraction_ring A K] section cyclotomic_ring /-- If `K` is the fraction field of `A`, the `A`-algebra structure on `cyclotomic_field n K`. -This is not an instance since it causes diamonds when `A = ℤ`. -/ +-/ @[nolint unused_arguments] -def cyclotomic_field.algebra_base : algebra A (cyclotomic_field n K) := -((algebra_map K (cyclotomic_field n K)).comp (algebra_map A K)).to_algebra +instance cyclotomic_field.algebra_base : algebra A (cyclotomic_field n K) := +splitting_field.algebra' (cyclotomic n K) -local attribute [instance] cyclotomic_field.algebra_base +/-- Ensure there are no diamonds when `A = ℤ`. -/ +example : algebra_int (cyclotomic_field n ℚ) = cyclotomic_field.algebra_base _ _ _ := rfl instance cyclotomic_field.no_zero_smul_divisors : no_zero_smul_divisors A (cyclotomic_field n K) := no_zero_smul_divisors.of_algebra_map_injective $ function.injective.comp diff --git a/src/number_theory/cyclotomic/rat.lean b/src/number_theory/cyclotomic/rat.lean index b7f1fdd822374..6481a65641071 100644 --- a/src/number_theory/cyclotomic/rat.lean +++ b/src/number_theory/cyclotomic/rat.lean @@ -129,32 +129,20 @@ begin exactI is_integral_closure_adjoin_singleton_of_prime_pow hζ, end -local attribute [-instance] cyclotomic_field.algebra - /-- 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 : char_zero ℚ := strict_ordered_semiring.to_char_zero, - haveI : is_cyclotomic_extension {p ^ k} ℚ (cyclotomic_field (p ^ k) ℚ), - { convert cyclotomic_field.is_cyclotomic_extension (p ^ k) _, - { exact subsingleton.elim _ _ }, - { exact ne_zero.char_zero } }, have hζ := zeta_spec (p ^ k) ℚ (cyclotomic_field (p ^ k) ℚ), refine ⟨is_fraction_ring.injective _ _, λ x, ⟨λ h, ⟨⟨x, _⟩, rfl⟩, _⟩⟩, { have := (is_integral_closure_adjoin_singleton_of_prime_pow hζ).is_integral_iff, obtain ⟨y, rfl⟩ := this.1 h, - convert adjoin_mono _ y.2, - { simp only [eq_iff_true_of_subsingleton] }, - { 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⟩, + refine adjoin_mono _ y.2, + simp only [pnat.pow_coe, set.singleton_subset_iff, set.mem_set_of_eq], + exact hζ.pow_eq_one }, + { rintro ⟨y, rfl⟩, exact is_integral.algebra_map ((is_cyclotomic_extension.integral {p ^ k} ℤ _) _) } end From 55ec6e9af7d3e0043f57e394cb06a72f6275273e Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Mon, 22 May 2023 14:39:43 +0000 Subject: [PATCH 1054/1291] chore(analysis/special_functions/non_integrable): split (#19047) Split out a fairly niche special-functions lemma, reducing the dependencies of this file. This decreases by 5 the length of the path to the Cauchy integral formula. --- src/analysis/special_functions/non_integrable.lean | 9 +++++---- 1 file changed, 5 insertions(+), 4 deletions(-) diff --git a/src/analysis/special_functions/non_integrable.lean b/src/analysis/special_functions/non_integrable.lean index d3ae8b75964d5..387f073c9f762 100644 --- a/src/analysis/special_functions/non_integrable.lean +++ b/src/analysis/special_functions/non_integrable.lean @@ -3,7 +3,8 @@ Copyright (c) 2021 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ -import analysis.special_functions.integrals +import analysis.special_functions.log.deriv +import measure_theory.integral.fund_thm_calculus /-! # Non integrable functions @@ -156,9 +157,9 @@ begin { refine λ h, or_iff_not_imp_left.2 (λ hne hc, _), exact not_interval_integrable_of_sub_inv_is_O_punctured (is_O_refl _ _) hne hc h }, { rintro (rfl|h₀), - exacts [interval_integrable.refl, - interval_integrable_inv (λ x hx, sub_ne_zero.2 $ ne_of_mem_of_not_mem hx h₀) - (continuous_on_id.sub continuous_on_const)] } + { exact interval_integrable.refl }, + refine ((continuous_sub_right c).continuous_on.inv₀ _).interval_integrable, + exact λ x hx, sub_ne_zero.2 $ ne_of_mem_of_not_mem hx h₀ } end /-- The function `λ x, x⁻¹` is integrable on `a..b` if and only if `a = b` or `0 ∉ [a, b]`. -/ From 75e7fca56381d056096ce5d05e938f63a6567828 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 22 May 2023 17:11:56 +0000 Subject: [PATCH 1055/1291] feat(analysis/calculus/mean_value): functions are equal if their derivatives and a point are equal (#19059) And the version for equality within a convex set. Co-authored-by: ADedecker --- src/analysis/calculus/mean_value.lean | 25 ++++++++++++++++++++++++- 1 file changed, 24 insertions(+), 1 deletion(-) diff --git a/src/analysis/calculus/mean_value.lean b/src/analysis/calculus/mean_value.lean index 4238fa2836b40..959f372245ec1 100644 --- a/src/analysis/calculus/mean_value.lean +++ b/src/analysis/calculus/mean_value.lean @@ -477,7 +477,7 @@ variables {𝕜 G : Type*} [is_R_or_C 𝕜] [normed_space 𝕜 E] [normed_add_co namespace convex -variables {f : E → G} {C : ℝ} {s : set E} {x y : E} {f' : E → E →L[𝕜] G} {φ : E →L[𝕜] G} +variables {f g : E → G} {C : ℝ} {s : set E} {x y : E} {f' g' : E → E →L[𝕜] G} {φ : E →L[𝕜] G} /-- The mean value theorem on a convex set: if the derivative of a function is bounded by `C`, then the function is `C`-Lipschitz. Version with `has_fderiv_within`. -/ @@ -638,6 +638,29 @@ theorem _root_.is_const_of_fderiv_eq_zero (hf : differentiable 𝕜 f) (hf' : convex_univ.is_const_of_fderiv_within_eq_zero hf.differentiable_on (λ x _, by rw fderiv_within_univ; exact hf' x) trivial trivial +/-- If two functions have equal Fréchet derivatives at every point of a convex set, and are equal at +one point in that set, then they are equal on that set. -/ +theorem eq_on_of_fderiv_within_eq (hs : convex ℝ s) + (hf : differentiable_on 𝕜 f s) (hg : differentiable_on 𝕜 g s) (hs' : unique_diff_on 𝕜 s) + (hf' : ∀ x ∈ s, fderiv_within 𝕜 f s x = fderiv_within 𝕜 g s x) (hx : x ∈ s) (hfgx : f x = g x) : + s.eq_on f g := +begin + intros y hy, + suffices : f x - g x = f y - g y, + { rwa [hfgx, sub_self, eq_comm, sub_eq_zero] at this }, + refine hs.is_const_of_fderiv_within_eq_zero (hf.sub hg) _ hx hy, + intros z hz, + rw [fderiv_within_sub (hs' _ hz) (hf _ hz) (hg _ hz), sub_eq_zero, hf' _ hz], +end + +theorem _root_.eq_of_fderiv_eq (hf : differentiable 𝕜 f) (hg : differentiable 𝕜 g) + (hf' : ∀ x, fderiv 𝕜 f x = fderiv 𝕜 g x) + (x : E) (hfgx : f x = g x) : + f = g := +suffices set.univ.eq_on f g, from funext $ λ x, this $ mem_univ x, +convex_univ.eq_on_of_fderiv_within_eq hf.differentiable_on hg.differentiable_on + unique_diff_on_univ (λ x hx, by simpa using hf' _) (mem_univ _) hfgx + end convex namespace convex From 6f8ab7de1c4b78a68ab8cf7dd83d549eb78a68a1 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 22 May 2023 22:49:48 +0000 Subject: [PATCH 1056/1291] feat(combinatorics/set_family/compression/uv): UV-compressing reduces the size of the shadow (#13149) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Prove `(∂ (𝓒 u v 𝒜)).card ≤ (∂ 𝒜).card`, which is key to Kruskal-Katona. Co-authored-by: Bhavik Mehta Co-authored-by: Mauricio Collares Co-authored-by: Eric Rodriguez --- .../set_family/compression/uv.lean | 197 ++++++++++++++++-- 1 file changed, 182 insertions(+), 15 deletions(-) diff --git a/src/combinatorics/set_family/compression/uv.lean b/src/combinatorics/set_family/compression/uv.lean index cec8c3800f4bb..704e6b0f708dc 100644 --- a/src/combinatorics/set_family/compression/uv.lean +++ b/src/combinatorics/set_family/compression/uv.lean @@ -3,7 +3,8 @@ Copyright (c) 2021 Bhavik Mehta. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies, Bhavik Mehta -/ -import data.finset.card +import combinatorics.set_family.shadow +import data.finset.sort /-! # UV-compressions @@ -27,7 +28,9 @@ minimise the shadow. It is the compressions of the elements of `s` whose compression is not already in `s` along with the element whose compression is already in `s`. This way of splitting into what moves and what does not ensures the compression doesn't squash the set family, which is proved by - `uv.card_compress`. + `uv.card_compression`. +* `uv.card_shadow_compression_le`: Compressing reduces the size of the shadow. This is a key fact in + the proof of Kruskal-Katona. ## Notation @@ -38,11 +41,6 @@ minimise the shadow. Even though our emphasis is on `finset α`, we define UV-compressions more generally in a generalized boolean algebra, so that one can use it for `set α`. -## TODO - -Prove that compressing reduces the size of shadow. This result and some more already exist on the -branch `combinatorics`. - ## References * https://github.com/b-mehta/maths-notes/blob/master/iii/mich/combinatorics.pdf @@ -79,8 +77,9 @@ variables [generalized_boolean_algebra α] [decidable_rel (@disjoint α _ _)] local attribute [instance] decidable_eq_of_decidable_le -/-- To UV-compress `a`, if it doesn't touch `U` and does contain `V`, we remove `V` and -put `U` in. We'll only really use this when `|U| = |V|` and `U ∩ V = ∅`. -/ +/-- UV-compressing `a` means removing `v` from it and adding `u` if `a` and `u` are disjoint and +`v ≤ a` (it replaces the `v` part of `a` by the `u` part). Else, UV-compressing `a` doesn't do +anything. This is most useful when `u` and `v` are disjoint finsets of the same size. -/ def compress (u v a : α) : α := if disjoint u a ∧ v ≤ a then (a ⊔ u) \ v else a /-- To UV-compress a set family, we compress each of its elements, except that we don't want to @@ -97,12 +96,20 @@ lemma compress_of_disjoint_of_le (hua : disjoint u a) (hva : v ≤ a) : compress u v a = (a ⊔ u) \ v := if_pos ⟨hua, hva⟩ +lemma compress_of_disjoint_of_le' (hva : disjoint v a) (hua : u ≤ a) : + compress u v ((a ⊔ v) \ u) = a := +by rw [compress_of_disjoint_of_le disjoint_sdiff_self_right + (le_sdiff.2 ⟨(le_sup_right : v ≤ a ⊔ v), hva.mono_right hua⟩), + sdiff_sup_cancel (le_sup_of_le_left hua), hva.symm.sup_sdiff_cancel_right] + /-- `a` is in the UV-compressed family iff it's in the original and its compression is in the original, or it's not in the original but it's the compression of something in the original. -/ lemma mem_compression : a ∈ 𝓒 u v s ↔ a ∈ s ∧ compress u v a ∈ s ∨ a ∉ s ∧ ∃ b ∈ s, compress u v b = a := by simp_rw [compression, mem_union, mem_filter, mem_image, and_comm (a ∉ s)] +protected lemma is_compressed.eq (h : is_compressed u v s) : 𝓒 u v s = s := h + @[simp] lemma compress_self (u a : α) : compress u u a = a := begin unfold compress, @@ -126,6 +133,14 @@ end /-- Any family is compressed along two identical elements. -/ lemma is_compressed_self (u : α) (s : finset α) : is_compressed u u s := compression_self u s +/-- An element can be compressed to any other element by removing/adding the differences. -/ +@[simp] lemma compress_sdiff_sdiff (a b : α) : compress (a \ b) (b \ a) b = a := +begin + refine (compress_of_disjoint_of_le disjoint_sdiff_self_left sdiff_le).trans _, + rw [sup_sdiff_self_right, sup_sdiff, disjoint_sdiff_self_right.sdiff_eq_left, sup_eq_right], + exact sdiff_sdiff_le, +end + lemma compress_disjoint (u v : α) : disjoint (s.filter (λ a, compress u v a ∈ s)) ((s.image $ compress u v).filter (λ a, a ∉ s)) := disjoint_left.2 $ λ a ha₁ ha₂, (mem_filter.1 ha₂).2 (mem_filter.1 ha₁).1 @@ -169,7 +184,7 @@ begin end /-- Compressing a family doesn't change its size. -/ -lemma card_compression (u v : α) (s : finset α) : (𝓒 u v s).card = s.card := +@[simp] lemma card_compression (u v : α) (s : finset α) : (𝓒 u v s).card = s.card := begin rw [compression, card_disjoint_union (compress_disjoint _ _), image_filter, card_image_of_inj_on, ←card_disjoint_union, filter_union_filter_neg_eq], @@ -187,6 +202,43 @@ begin { exact (ha.2 ha.1).elim } end +lemma le_of_mem_compression_of_not_mem (h : a ∈ 𝓒 u v s) (ha : a ∉ s) : u ≤ a := +begin + rw mem_compression at h, + obtain _ | ⟨-, b, hb, hba⟩ := h, + { cases ha h.1 }, + unfold compress at hba, + split_ifs at hba, + { rw [←hba, le_sdiff], + exact ⟨le_sup_right, h.1.mono_right h.2⟩ }, + { cases ne_of_mem_of_not_mem hb ha hba } +end + +lemma disjoint_of_mem_compression_of_not_mem (h : a ∈ 𝓒 u v s) (ha : a ∉ s) : disjoint v a := +begin + rw mem_compression at h, + obtain _ | ⟨-, b, hb, hba⟩ := h, + { cases ha h.1 }, + unfold compress at hba, + split_ifs at hba, + { rw ←hba, + exact disjoint_sdiff_self_right }, + { cases ne_of_mem_of_not_mem hb ha hba } +end + +lemma sup_sdiff_mem_of_mem_compression_of_not_mem (h : a ∈ 𝓒 u v s) (ha : a ∉ s) : + (a ⊔ v) \ u ∈ s := +begin + rw mem_compression at h, + obtain _ | ⟨-, b, hb, hba⟩ := h, + { cases ha h.1 }, + unfold compress at hba, + split_ifs at hba, + { rwa [←hba, sdiff_sup_cancel (le_sup_of_le_left h.2), sup_sdiff_right_self, + h.1.symm.sdiff_eq_left] }, + { cases ne_of_mem_of_not_mem hb ha hba } +end + /-- If `a` is in the family compression and can be compressed, then its compression is in the original family. -/ lemma sup_sdiff_mem_of_mem_compression (ha : a ∈ 𝓒 u v s) (hva : v ≤ a) (hua : disjoint u a) : @@ -219,8 +271,7 @@ begin unfold compress at h, split_ifs at h, { rw [←h, le_sdiff_iff] at hva, - rw [hvu hva, hva, sup_bot_eq, sdiff_bot] at h, - rwa ←h }, + rwa [←h, hvu hva, hva, sup_bot_eq, sdiff_bot] }, { rwa ←h } end @@ -230,16 +281,132 @@ end generalized_boolean_algebra open_locale finset_family -variables [decidable_eq α] {𝒜 : finset (finset α)} {U V A : finset α} +variables [decidable_eq α] {𝒜 : finset (finset α)} {u v a : finset α} /-- Compressing a finset doesn't change its size. -/ -lemma card_compress (hUV : U.card = V.card) (A : finset α) : (compress U V A).card = A.card := +lemma card_compress (hUV : u.card = v.card) (A : finset α) : (compress u v A).card = A.card := begin unfold compress, split_ifs, { rw [card_sdiff (h.2.trans le_sup_left), sup_eq_union, card_disjoint_union h.1.symm, hUV, - add_tsub_cancel_right] }, + add_tsub_cancel_right] }, { refl } end +private lemma aux (huv : ∀ x ∈ u, ∃ y ∈ v, is_compressed (u.erase x) (v.erase y) 𝒜) : + v = ∅ → u = ∅ := +by { rintro rfl, refine eq_empty_of_forall_not_mem (λ a ha, _), obtain ⟨_, ⟨⟩, -⟩ := huv a ha } + +/-- UV-compression reduces the size of the shadow of `𝒜` if, for all `x ∈ u` there is `y ∈ v` such +that `𝒜` is `(u.erase x, v.erase y)`-compressed. This is the key fact about compression for +Kruskal-Katona. -/ +lemma shadow_compression_subset_compression_shadow (u v : finset α) + (huv : ∀ x ∈ u, ∃ y ∈ v, is_compressed (u.erase x) (v.erase y) 𝒜) : + ∂ (𝓒 u v 𝒜) ⊆ 𝓒 u v (∂ 𝒜) := +begin + set 𝒜' := 𝓒 u v 𝒜, + suffices H : ∀ s, s ∈ ∂ 𝒜' → s ∉ ∂ 𝒜 → + u ⊆ s ∧ disjoint v s ∧ (s ∪ v) \ u ∈ ∂ 𝒜 ∧ (s ∪ v) \ u ∉ ∂ 𝒜', + { rintro s hs', + rw mem_compression, + by_cases hs : s ∈ 𝒜.shadow, swap, + { obtain ⟨hus, hvs, h, _⟩ := H _ hs' hs, + exact or.inr ⟨hs, _, h, compress_of_disjoint_of_le' hvs hus⟩ }, + refine or.inl ⟨hs, _⟩, + rw compress, + split_ifs with huvs, swap, + { exact hs }, + rw mem_shadow_iff at hs', + obtain ⟨t, Ht, a, hat, rfl⟩ := hs', + have hav : a ∉ v := not_mem_mono huvs.2 (not_mem_erase a t), + have hvt : v ≤ t := huvs.2.trans (erase_subset _ t), + have ht : t ∈ 𝒜 := mem_of_mem_compression Ht hvt (aux huv), + by_cases hau : a ∈ u, + { obtain ⟨b, hbv, Hcomp⟩ := huv a hau, + refine mem_shadow_iff_insert_mem.2 ⟨b, not_mem_sdiff_of_mem_right hbv, _⟩, + rw ←Hcomp.eq at ht, + have hsb := sup_sdiff_mem_of_mem_compression ht ((erase_subset _ _).trans hvt) + (disjoint_erase_comm.2 huvs.1), + rwa [sup_eq_union, sdiff_erase (mem_union_left _ $ hvt hbv), union_erase_of_mem hat, + ←erase_union_of_mem hau] at hsb }, + { refine mem_shadow_iff.2 ⟨(t ⊔ u) \ v, sup_sdiff_mem_of_mem_compression Ht hvt $ + disjoint_of_erase_right hau huvs.1, a, _, _⟩, + { rw [sup_eq_union, mem_sdiff, mem_union], + exact ⟨or.inl hat, hav⟩ }, + { rw [←erase_sdiff_comm, sup_eq_union, erase_union_distrib, erase_eq_of_not_mem hau] } } }, + intros s hs𝒜' hs𝒜, + -- This is gonna be useful a couple of times so let's name it. + have m : ∀ y ∉ s, insert y s ∉ 𝒜 := λ y h a, hs𝒜 (mem_shadow_iff_insert_mem.2 ⟨y, h, a⟩), + obtain ⟨x, _, _⟩ := mem_shadow_iff_insert_mem.1 hs𝒜', + have hus : u ⊆ insert x s := le_of_mem_compression_of_not_mem ‹_ ∈ 𝒜'› (m _ ‹x ∉ s›), + have hvs : disjoint v (insert x s) := disjoint_of_mem_compression_of_not_mem ‹_› (m _ ‹x ∉ s›), + have : (insert x s ∪ v) \ u ∈ 𝒜 := sup_sdiff_mem_of_mem_compression_of_not_mem ‹_› (m _ ‹x ∉ s›), + have hsv : disjoint s v := hvs.symm.mono_left (subset_insert _ _), + have hvu : disjoint v u := disjoint_of_subset_right hus hvs, + have hxv : x ∉ v := disjoint_right.1 hvs (mem_insert_self _ _), + have : v \ u = v := ‹disjoint v u›.sdiff_eq_left, + -- The first key part is that `x ∉ u` + have : x ∉ u, + { intro hxu, + obtain ⟨y, hyv, hxy⟩ := huv x hxu, + -- If `x ∈ u`, we can get `y ∈ v` so that `𝒜` is `(u.erase x, v.erase y)`-compressed + apply m y (disjoint_right.1 hsv hyv), + -- and we will use this `y` to contradict `m`, so we would like to show `insert y s ∈ 𝒜`. + -- We do this by showing the below + have : ((insert x s ∪ v) \ u ∪ erase u x) \ erase v y ∈ 𝒜, + { refine sup_sdiff_mem_of_mem_compression (by rwa hxy.eq) _ + (disjoint_of_subset_left (erase_subset _ _) disjoint_sdiff), + rw [union_sdiff_distrib, ‹v \ u = v›], + exact (erase_subset _ _).trans (subset_union_right _ _) }, + -- and then arguing that it's the same + convert this, + rw [sdiff_union_erase_cancel (hus.trans $ subset_union_left _ _) ‹x ∈ u›, erase_union_distrib, + erase_insert ‹x ∉ s›, erase_eq_of_not_mem ‹x ∉ v›, sdiff_erase (mem_union_right _ hyv), + union_sdiff_cancel_right hsv] }, + -- Now that this is done, it's immediate that `u ⊆ s` + have hus : u ⊆ s, + { rwa [←erase_eq_of_not_mem ‹x ∉ u›, ←subset_insert_iff] }, + -- and we already had that `v` and `s` are disjoint, + -- so it only remains to get `(s ∪ v) \ u ∈ ∂ 𝒜 \ ∂ 𝒜'` + simp_rw [mem_shadow_iff_insert_mem], + refine ⟨hus, hsv.symm, ⟨x, _, _⟩, _⟩, + -- `(s ∪ v) \ u ∈ ∂ 𝒜` is pretty direct: + { exact not_mem_sdiff_of_not_mem_left (not_mem_union.2 ⟨‹x ∉ s›, ‹x ∉ v›⟩) }, + { rwa [←insert_sdiff_of_not_mem _ ‹x ∉ u›, ←insert_union] }, + -- For (s ∪ v) \ u ∉ ∂ 𝒜', we split up based on w ∈ u + rintro ⟨w, hwB, hw𝒜'⟩, + have : v ⊆ insert w ((s ∪ v) \ u) := (subset_sdiff.2 ⟨subset_union_right _ _, hvu⟩).trans + (subset_insert _ _), + by_cases hwu : w ∈ u, + -- If `w ∈ u`, we find `z ∈ v`, and contradict `m` again + { obtain ⟨z, hz, hxy⟩ := huv w hwu, + apply m z (disjoint_right.1 hsv hz), + have : insert w ((s ∪ v) \ u) ∈ 𝒜 := mem_of_mem_compression hw𝒜' ‹_› (aux huv), + have : (insert w ((s ∪ v) \ u) ∪ erase u w) \ erase v z ∈ 𝒜, + { refine sup_sdiff_mem_of_mem_compression (by rwa hxy.eq) ((erase_subset _ _).trans ‹_›) _, + rw ←sdiff_erase (mem_union_left _ $ hus hwu), + exact disjoint_sdiff }, + convert this, + rw [insert_union_comm, insert_erase ‹w ∈ u›, sdiff_union_of_subset + (hus.trans $ subset_union_left _ _), sdiff_erase (mem_union_right _ ‹z ∈ v›), + union_sdiff_cancel_right hsv] }, + -- If `w ∉ u`, we contradict `m` again + rw [mem_sdiff, ←not_imp, not_not] at hwB, + apply m w (hwu ∘ hwB ∘ mem_union_left _), + have : (insert w ((s ∪ v) \ u) ∪ u) \ v ∈ 𝒜 := sup_sdiff_mem_of_mem_compression + ‹insert w ((s ∪ v) \ u) ∈ 𝒜'› ‹_› (disjoint_insert_right.2 ⟨‹_›, disjoint_sdiff⟩), + convert this, + rw [insert_union, sdiff_union_of_subset (hus.trans $ subset_union_left _ _), + insert_sdiff_of_not_mem _ (hwu ∘ hwB ∘ mem_union_right _), union_sdiff_cancel_right hsv], +end + +/-- UV-compression reduces the size of the shadow of `𝒜` if, for all `x ∈ u` there is `y ∈ v` +such that `𝒜` is `(u.erase x, v.erase y)`-compressed. This is the key UV-compression fact needed for +Kruskal-Katona. -/ +lemma card_shadow_compression_le (u v : finset α) + (huv : ∀ x ∈ u, ∃ y ∈ v, is_compressed (u.erase x) (v.erase y) 𝒜) : + (∂ (𝓒 u v 𝒜)).card ≤ (∂ 𝒜).card := +(card_le_of_subset $ shadow_compression_subset_compression_shadow _ _ huv).trans + (card_compression _ _ _).le + end uv From 38df578a6450a8c5142b3727e3ae894c2300cae0 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Tue, 23 May 2023 01:09:14 +0000 Subject: [PATCH 1057/1291] chore(*): add mathlib4 synchronization comments (#19063) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.algebraic_card` * `algebra.category.GroupWithZero` * `algebra.category.Ring.instances` * `algebra.dual_quaternion` * `algebra.quaternion` * `algebra.quaternion_basis` * `analysis.calculus.fderiv.add` * `analysis.calculus.fderiv.comp` * `analysis.calculus.fderiv.equiv` * `analysis.calculus.fderiv.linear` * `analysis.calculus.fderiv.prod` * `analysis.calculus.fderiv.restrict_scalars` * `analysis.convex.between` * `analysis.convex.strict_convex_between` * `analysis.locally_convex.strong_topology` * `combinatorics.configuration` * `data.real.irrational` * `dynamics.ergodic.conservative` * `field_theory.minpoly.basic` * `field_theory.minpoly.field` * `group_theory.perm.cycle.concrete` * `linear_algebra.bilinear_form` * `linear_algebra.bilinear_form.tensor_product` * `linear_algebra.coevaluation` * `linear_algebra.contraction` * `linear_algebra.dual` * `linear_algebra.matrix.dual` * `measure_theory.constructions.borel_space.metrizable` * `measure_theory.covering.vitali` * `measure_theory.function.special_functions.is_R_or_C` * `measure_theory.integral.lebesgue_normed_space` * `measure_theory.integral.riesz_markov_kakutani` * `measure_theory.measure.content` * `measure_theory.measure.giry_monad` * `number_theory.legendre_symbol.zmod_char` * `number_theory.padics.padic_numbers` * `order.interval` * `ring_theory.algebraic` * `ring_theory.artinian` * `ring_theory.discrete_valuation_ring.basic` * `ring_theory.graded_algebra.radical` * `ring_theory.integral_closure` * `ring_theory.localization.as_subring` * `ring_theory.localization.away.basic` * `ring_theory.localization.cardinality` * `ring_theory.localization.localization_localization` * `ring_theory.power_series.basic` * `ring_theory.power_series.well_known` * `ring_theory.witt_vector.witt_polynomial` * `topology.metric_space.metrizable` * `topology.metric_space.thickened_indicator` * `topology.sheaves.punit` * `topology.urysohns_bounded` * `topology.vector_bundle.basic` * `topology.vector_bundle.constructions` --- src/algebra/algebraic_card.lean | 3 +++ src/algebra/category/GroupWithZero.lean | 3 +++ src/algebra/category/Ring/instances.lean | 3 +++ src/algebra/dual_quaternion.lean | 3 +++ src/algebra/quaternion.lean | 3 +++ src/algebra/quaternion_basis.lean | 3 +++ src/analysis/calculus/fderiv/add.lean | 3 +++ src/analysis/calculus/fderiv/comp.lean | 3 +++ src/analysis/calculus/fderiv/equiv.lean | 3 +++ src/analysis/calculus/fderiv/linear.lean | 3 +++ src/analysis/calculus/fderiv/prod.lean | 3 +++ src/analysis/calculus/fderiv/restrict_scalars.lean | 3 +++ src/analysis/convex/between.lean | 3 +++ src/analysis/convex/strict_convex_between.lean | 3 +++ src/analysis/locally_convex/strong_topology.lean | 3 +++ src/combinatorics/configuration.lean | 3 +++ src/data/real/irrational.lean | 3 +++ src/dynamics/ergodic/conservative.lean | 3 +++ src/field_theory/minpoly/basic.lean | 3 +++ src/field_theory/minpoly/field.lean | 3 +++ src/group_theory/perm/cycle/concrete.lean | 3 +++ src/linear_algebra/bilinear_form.lean | 3 +++ src/linear_algebra/bilinear_form/tensor_product.lean | 3 +++ src/linear_algebra/coevaluation.lean | 3 +++ src/linear_algebra/contraction.lean | 3 +++ src/linear_algebra/dual.lean | 3 +++ src/linear_algebra/matrix/dual.lean | 3 +++ src/measure_theory/constructions/borel_space/metrizable.lean | 3 +++ src/measure_theory/covering/vitali.lean | 3 +++ src/measure_theory/function/special_functions/is_R_or_C.lean | 3 +++ src/measure_theory/integral/lebesgue_normed_space.lean | 5 ++++- src/measure_theory/integral/riesz_markov_kakutani.lean | 3 +++ src/measure_theory/measure/content.lean | 3 +++ src/measure_theory/measure/giry_monad.lean | 3 +++ src/number_theory/legendre_symbol/zmod_char.lean | 3 +++ src/number_theory/padics/padic_numbers.lean | 3 +++ src/order/interval.lean | 3 +++ src/ring_theory/algebraic.lean | 3 +++ src/ring_theory/artinian.lean | 3 +++ src/ring_theory/discrete_valuation_ring/basic.lean | 3 +++ src/ring_theory/graded_algebra/radical.lean | 3 +++ src/ring_theory/integral_closure.lean | 3 +++ src/ring_theory/localization/as_subring.lean | 3 +++ src/ring_theory/localization/away/basic.lean | 3 +++ src/ring_theory/localization/cardinality.lean | 3 +++ src/ring_theory/localization/localization_localization.lean | 3 +++ src/ring_theory/power_series/basic.lean | 3 +++ src/ring_theory/power_series/well_known.lean | 3 +++ src/ring_theory/witt_vector/witt_polynomial.lean | 3 +++ src/topology/metric_space/metrizable.lean | 3 +++ src/topology/metric_space/thickened_indicator.lean | 3 +++ src/topology/sheaves/punit.lean | 3 +++ src/topology/urysohns_bounded.lean | 3 +++ src/topology/vector_bundle/basic.lean | 3 +++ src/topology/vector_bundle/constructions.lean | 3 +++ 55 files changed, 166 insertions(+), 1 deletion(-) diff --git a/src/algebra/algebraic_card.lean b/src/algebra/algebraic_card.lean index 5cf00a5d2af68..1d88761edeca9 100644 --- a/src/algebra/algebraic_card.lean +++ b/src/algebra/algebraic_card.lean @@ -10,6 +10,9 @@ import ring_theory.algebraic /-! ### Cardinality of algebraic numbers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we prove variants of the following result: the cardinality of algebraic numbers under an R-algebra is at most `# R[X] * ℵ₀`. diff --git a/src/algebra/category/GroupWithZero.lean b/src/algebra/category/GroupWithZero.lean index f97a34ef4733c..522d953a2e396 100644 --- a/src/algebra/category/GroupWithZero.lean +++ b/src/algebra/category/GroupWithZero.lean @@ -9,6 +9,9 @@ import algebra.category.Mon.basic /-! # The category of groups with zero +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines `GroupWithZero`, the category of groups with zero. -/ diff --git a/src/algebra/category/Ring/instances.lean b/src/algebra/category/Ring/instances.lean index cef4e306a2150..009b23397c567 100644 --- a/src/algebra/category/Ring/instances.lean +++ b/src/algebra/category/Ring/instances.lean @@ -9,6 +9,9 @@ import ring_theory.ideal.local_ring /-! # Ring-theoretic results in terms of categorical languages + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ open category_theory diff --git a/src/algebra/dual_quaternion.lean b/src/algebra/dual_quaternion.lean index 863840d77d885..2937d89336bf2 100644 --- a/src/algebra/dual_quaternion.lean +++ b/src/algebra/dual_quaternion.lean @@ -9,6 +9,9 @@ import algebra.quaternion /-! # Dual quaternions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Similar to the way that rotations in 3D space can be represented by quaternions of unit length, rigid motions in 3D space can be represented by dual quaternions of unit length. diff --git a/src/algebra/quaternion.lean b/src/algebra/quaternion.lean index c465eb0e0fb68..449dbc5ee7f43 100644 --- a/src/algebra/quaternion.lean +++ b/src/algebra/quaternion.lean @@ -13,6 +13,9 @@ import tactic.ring_exp /-! # Quaternions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define quaternions `ℍ[R]` over a commutative ring `R`, and define some algebraic structures on `ℍ[R]`. diff --git a/src/algebra/quaternion_basis.lean b/src/algebra/quaternion_basis.lean index d65824650cc52..deeee0cd9c997 100644 --- a/src/algebra/quaternion_basis.lean +++ b/src/algebra/quaternion_basis.lean @@ -9,6 +9,9 @@ import tactic.ring /-! # Basis on a quaternion-like algebra +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions * `quaternion_algebra.basis A c₁ c₂`: a basis for a subspace of an `R`-algebra `A` that has the diff --git a/src/analysis/calculus/fderiv/add.lean b/src/analysis/calculus/fderiv/add.lean index 98c66b263e07c..459bb39412abe 100644 --- a/src/analysis/calculus/fderiv/add.lean +++ b/src/analysis/calculus/fderiv/add.lean @@ -9,6 +9,9 @@ import analysis.calculus.fderiv.comp /-! # Additive operations on derivatives +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + For detailed documentation of the Fréchet derivative, see the module docstring of `analysis/calculus/fderiv/basic.lean`. diff --git a/src/analysis/calculus/fderiv/comp.lean b/src/analysis/calculus/fderiv/comp.lean index 9304da4ab33cd..4efc624e08dee 100644 --- a/src/analysis/calculus/fderiv/comp.lean +++ b/src/analysis/calculus/fderiv/comp.lean @@ -8,6 +8,9 @@ import analysis.calculus.fderiv.basic /-! # The derivative of a composition (chain rule) +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + For detailed documentation of the Fréchet derivative, see the module docstring of `analysis/calculus/fderiv/basic.lean`. diff --git a/src/analysis/calculus/fderiv/equiv.lean b/src/analysis/calculus/fderiv/equiv.lean index 1795a6b6494a0..b3fdd1441a863 100644 --- a/src/analysis/calculus/fderiv/equiv.lean +++ b/src/analysis/calculus/fderiv/equiv.lean @@ -9,6 +9,9 @@ import analysis.calculus.fderiv.comp /-! # The derivative of a linear equivalence +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + For detailed documentation of the Fréchet derivative, see the module docstring of `analysis/calculus/fderiv/basic.lean`. diff --git a/src/analysis/calculus/fderiv/linear.lean b/src/analysis/calculus/fderiv/linear.lean index 28b63c073871e..73102958b1d4c 100644 --- a/src/analysis/calculus/fderiv/linear.lean +++ b/src/analysis/calculus/fderiv/linear.lean @@ -8,6 +8,9 @@ import analysis.calculus.fderiv.basic /-! # The derivative of bounded linear maps +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + For detailed documentation of the Fréchet derivative, see the module docstring of `analysis/calculus/fderiv/basic.lean`. diff --git a/src/analysis/calculus/fderiv/prod.lean b/src/analysis/calculus/fderiv/prod.lean index 01b3a4f00d183..007fba3bc694d 100644 --- a/src/analysis/calculus/fderiv/prod.lean +++ b/src/analysis/calculus/fderiv/prod.lean @@ -9,6 +9,9 @@ import analysis.calculus.fderiv.comp /-! # Derivative of the cartesian product of functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + For detailed documentation of the Fréchet derivative, see the module docstring of `analysis/calculus/fderiv/basic.lean`. diff --git a/src/analysis/calculus/fderiv/restrict_scalars.lean b/src/analysis/calculus/fderiv/restrict_scalars.lean index eddc2555b61a3..d414c3f156487 100644 --- a/src/analysis/calculus/fderiv/restrict_scalars.lean +++ b/src/analysis/calculus/fderiv/restrict_scalars.lean @@ -8,6 +8,9 @@ import analysis.calculus.fderiv.basic /-! # The derivative of the scalar restriction of a linear map +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + For detailed documentation of the Fréchet derivative, see the module docstring of `analysis/calculus/fderiv/basic.lean`. diff --git a/src/analysis/convex/between.lean b/src/analysis/convex/between.lean index 6c1c1e94a8834..2679b932ce622 100644 --- a/src/analysis/convex/between.lean +++ b/src/analysis/convex/between.lean @@ -12,6 +12,9 @@ import algebra.char_p.invertible /-! # Betweenness in affine spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines notions of a point in an affine space being between two given points. ## Main definitions diff --git a/src/analysis/convex/strict_convex_between.lean b/src/analysis/convex/strict_convex_between.lean index 2e32dcbc204d5..0cfcd2f643848 100644 --- a/src/analysis/convex/strict_convex_between.lean +++ b/src/analysis/convex/strict_convex_between.lean @@ -9,6 +9,9 @@ import analysis.convex.strict_convex_space /-! # Betweenness in affine spaces for strictly convex spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves results about betweenness for points in an affine space for a strictly convex space. diff --git a/src/analysis/locally_convex/strong_topology.lean b/src/analysis/locally_convex/strong_topology.lean index a5e444a5729da..24943723ea67b 100644 --- a/src/analysis/locally_convex/strong_topology.lean +++ b/src/analysis/locally_convex/strong_topology.lean @@ -9,6 +9,9 @@ import topology.algebra.module.locally_convex /-! # Local convexity of the strong topology +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove that the strong topology on `E →L[ℝ] F` is locally convex provided that `F` is locally convex. diff --git a/src/combinatorics/configuration.lean b/src/combinatorics/configuration.lean index fdbb9de545df6..28b4ea67e48e5 100644 --- a/src/combinatorics/configuration.lean +++ b/src/combinatorics/configuration.lean @@ -10,6 +10,9 @@ import set_theory.cardinal.finite /-! # Configurations of Points and lines + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. This file introduces abstract configurations of points and lines, and proves some basic properties. ## Main definitions diff --git a/src/data/real/irrational.lean b/src/data/real/irrational.lean index c2463026fe019..fadcdd66a722a 100644 --- a/src/data/real/irrational.lean +++ b/src/data/real/irrational.lean @@ -11,6 +11,9 @@ import ring_theory.int.basic /-! # Irrational real numbers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define a predicate `irrational` on `ℝ`, prove that the `n`-th root of an integer number is irrational if it is not integer, and that `sqrt q` is irrational if and only if `rat.sqrt q * rat.sqrt q ≠ q ∧ 0 ≤ q`. diff --git a/src/dynamics/ergodic/conservative.lean b/src/dynamics/ergodic/conservative.lean index c2a2fb13233c4..19c1596937165 100644 --- a/src/dynamics/ergodic/conservative.lean +++ b/src/dynamics/ergodic/conservative.lean @@ -10,6 +10,9 @@ import combinatorics.pigeonhole /-! # Conservative systems +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define `f : α → α` to be a *conservative* system w.r.t a measure `μ` if `f` is non-singular (`measure_theory.quasi_measure_preserving`) and for every measurable set `s` of positive measure at least one point `x ∈ s` returns back to `s` after some number of iterations of diff --git a/src/field_theory/minpoly/basic.lean b/src/field_theory/minpoly/basic.lean index 121a0388520f2..4a19907b5e231 100644 --- a/src/field_theory/minpoly/basic.lean +++ b/src/field_theory/minpoly/basic.lean @@ -8,6 +8,9 @@ import ring_theory.integral_closure /-! # Minimal polynomials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the minimal polynomial of an element `x` of an `A`-algebra `B`, under the assumption that x is integral over `A`, and derives some basic properties such as ireducibility under the assumption `B` is a domain. diff --git a/src/field_theory/minpoly/field.lean b/src/field_theory/minpoly/field.lean index 4bfcbc603bd40..dc6c8381ecff9 100644 --- a/src/field_theory/minpoly/field.lean +++ b/src/field_theory/minpoly/field.lean @@ -10,6 +10,9 @@ import ring_theory.algebraic /-! # Minimal polynomials on an algebra over a field +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file specializes the theory of minpoly to the setting of field extensions and derives some well-known properties, amongst which the fact that minimal polynomials are irreducible, and uniquely determined by their defining property. diff --git a/src/group_theory/perm/cycle/concrete.lean b/src/group_theory/perm/cycle/concrete.lean index d52b8721d9ebc..8d4f2ac2b3d63 100644 --- a/src/group_theory/perm/cycle/concrete.lean +++ b/src/group_theory/perm/cycle/concrete.lean @@ -11,6 +11,9 @@ import group_theory.perm.list # Properties of cyclic permutations constructed from lists/cycles +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In the following, `{α : Type*} [fintype α] [decidable_eq α]`. ## Main definitions diff --git a/src/linear_algebra/bilinear_form.lean b/src/linear_algebra/bilinear_form.lean index 8befd53e09164..c2aff4436f65f 100644 --- a/src/linear_algebra/bilinear_form.lean +++ b/src/linear_algebra/bilinear_form.lean @@ -10,6 +10,9 @@ import linear_algebra.free_module.finite.matrix /-! # Bilinear form +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines a bilinear form over a module. Basic ideas such as orthogonality are also introduced, as well as reflexivive, symmetric, non-degenerate and alternating bilinear forms. Adjoints of diff --git a/src/linear_algebra/bilinear_form/tensor_product.lean b/src/linear_algebra/bilinear_form/tensor_product.lean index 97fdc9fcd996f..47f45b2f6360a 100644 --- a/src/linear_algebra/bilinear_form/tensor_product.lean +++ b/src/linear_algebra/bilinear_form/tensor_product.lean @@ -9,6 +9,9 @@ import linear_algebra.tensor_product /-! # The bilinear form on a tensor product +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions * `bilin_form.tensor_distrib (B₁ ⊗ₜ B₂)`: the bilinear form on `M₁ ⊗ M₂` constructed by applying diff --git a/src/linear_algebra/coevaluation.lean b/src/linear_algebra/coevaluation.lean index 32dcc3dd480c5..31abc9c6cc8e8 100644 --- a/src/linear_algebra/coevaluation.lean +++ b/src/linear_algebra/coevaluation.lean @@ -10,6 +10,9 @@ import linear_algebra.dual /-! # The coevaluation map on finite dimensional vector spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given a finite dimensional vector space `V` over a field `K` this describes the canonical linear map from `K` to `V ⊗ dual K V` which corresponds to the identity function on `V`. diff --git a/src/linear_algebra/contraction.lean b/src/linear_algebra/contraction.lean index 13f941f92f2f8..cbaa519592cdf 100644 --- a/src/linear_algebra/contraction.lean +++ b/src/linear_algebra/contraction.lean @@ -9,6 +9,9 @@ import linear_algebra.matrix.to_lin /-! # Contractions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given modules $M, N$ over a commutative ring $R$, this file defines the natural linear maps: $M^* \otimes M \to R$, $M \otimes M^* \to R$, and $M^* \otimes N → Hom(M, N)$, as well as proving some basic properties of these maps. diff --git a/src/linear_algebra/dual.lean b/src/linear_algebra/dual.lean index af61603118098..fc9421636e75a 100644 --- a/src/linear_algebra/dual.lean +++ b/src/linear_algebra/dual.lean @@ -12,6 +12,9 @@ import linear_algebra.free_module.finite.basic /-! # Dual vector spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The dual space of an $R$-module $M$ is the $R$-module of $R$-linear maps $M \to R$. ## Main definitions diff --git a/src/linear_algebra/matrix/dual.lean b/src/linear_algebra/matrix/dual.lean index 46fc198a39498..4d112dd3c56d3 100644 --- a/src/linear_algebra/matrix/dual.lean +++ b/src/linear_algebra/matrix/dual.lean @@ -8,6 +8,9 @@ import linear_algebra.matrix.to_lin /-! # Dual space, linear maps and matrices. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains some results on the matrix corresponding to the transpose of a linear map (in the dual space). diff --git a/src/measure_theory/constructions/borel_space/metrizable.lean b/src/measure_theory/constructions/borel_space/metrizable.lean index 6572f6f437674..36478e37f5677 100644 --- a/src/measure_theory/constructions/borel_space/metrizable.lean +++ b/src/measure_theory/constructions/borel_space/metrizable.lean @@ -8,6 +8,9 @@ import topology.metric_space.metrizable /-! # Measurable functions in (pseudo-)metrizable Borel spaces + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ open filter measure_theory topological_space diff --git a/src/measure_theory/covering/vitali.lean b/src/measure_theory/covering/vitali.lean index 622c70dee1868..00af34b0a68e1 100644 --- a/src/measure_theory/covering/vitali.lean +++ b/src/measure_theory/covering/vitali.lean @@ -10,6 +10,9 @@ import measure_theory.covering.vitali_family /-! # Vitali covering theorems +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The topological Vitali covering theorem, in its most classical version, states the following. Consider a family of balls `(B (x_i, r_i))_{i ∈ I}` in a metric space, with uniformly bounded radii. Then one can extract a disjoint subfamily indexed by `J ⊆ I`, such that any `B (x_i, r_i)` diff --git a/src/measure_theory/function/special_functions/is_R_or_C.lean b/src/measure_theory/function/special_functions/is_R_or_C.lean index 5ec43cd5ae171..105fbac18bb90 100644 --- a/src/measure_theory/function/special_functions/is_R_or_C.lean +++ b/src/measure_theory/function/special_functions/is_R_or_C.lean @@ -10,6 +10,9 @@ import data.is_R_or_C.lemmas /-! # Measurability of the basic `is_R_or_C` functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + -/ noncomputable theory diff --git a/src/measure_theory/integral/lebesgue_normed_space.lean b/src/measure_theory/integral/lebesgue_normed_space.lean index 4301c8df6ba13..04fc83750aa4b 100644 --- a/src/measure_theory/integral/lebesgue_normed_space.lean +++ b/src/measure_theory/integral/lebesgue_normed_space.lean @@ -6,7 +6,10 @@ Authors: Sébastien Gouëzel import measure_theory.integral.lebesgue import analysis.normed_space.basic -/-! # A lemma about measurability with density under scalar multiplication in normed spaces -/ +/-! # A lemma about measurability with density under scalar multiplication in normed spaces + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4.-/ open measure_theory filter ennreal set open_locale nnreal ennreal diff --git a/src/measure_theory/integral/riesz_markov_kakutani.lean b/src/measure_theory/integral/riesz_markov_kakutani.lean index aeff61a56f4f4..6b25ad04cfa05 100644 --- a/src/measure_theory/integral/riesz_markov_kakutani.lean +++ b/src/measure_theory/integral/riesz_markov_kakutani.lean @@ -9,6 +9,9 @@ import topology.sets.compacts /-! # Riesz–Markov–Kakutani representation theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file will prove different versions of the Riesz-Markov-Kakutani representation theorem. The theorem is first proven for compact spaces, from which the statements about linear functionals on bounded continuous functions or compactly supported functions on locally compact spaces follow. diff --git a/src/measure_theory/measure/content.lean b/src/measure_theory/measure/content.lean index 3ecbc9123f914..b7c6e0d7cff91 100644 --- a/src/measure_theory/measure/content.lean +++ b/src/measure_theory/measure/content.lean @@ -10,6 +10,9 @@ import topology.sets.compacts /-! # Contents +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we work with *contents*. A content `λ` is a function from a certain class of subsets (such as the compact subsets) to `ℝ≥0` that is * additive: If `K₁` and `K₂` are disjoint sets in the domain of `λ`, diff --git a/src/measure_theory/measure/giry_monad.lean b/src/measure_theory/measure/giry_monad.lean index 2439ee3104357..54df828191871 100644 --- a/src/measure_theory/measure/giry_monad.lean +++ b/src/measure_theory/measure/giry_monad.lean @@ -8,6 +8,9 @@ import measure_theory.integral.lebesgue /-! # The Giry monad +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Let X be a measurable space. The collection of all measures on X again forms a measurable space. This construction forms a monad on measurable spaces and measurable functions, called the Giry monad. diff --git a/src/number_theory/legendre_symbol/zmod_char.lean b/src/number_theory/legendre_symbol/zmod_char.lean index dde190d57fb85..5921d6bbd543d 100644 --- a/src/number_theory/legendre_symbol/zmod_char.lean +++ b/src/number_theory/legendre_symbol/zmod_char.lean @@ -10,6 +10,9 @@ import number_theory.legendre_symbol.mul_character /-! # Quadratic characters on ℤ/nℤ +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines some quadratic characters on the rings ℤ/4ℤ and ℤ/8ℤ. We set them up to be of type `mul_char (zmod n) ℤ`, where `n` is `4` or `8`. diff --git a/src/number_theory/padics/padic_numbers.lean b/src/number_theory/padics/padic_numbers.lean index fdd6b9baf34ce..778e8937e1c6e 100644 --- a/src/number_theory/padics/padic_numbers.lean +++ b/src/number_theory/padics/padic_numbers.lean @@ -9,6 +9,9 @@ import analysis.normed.field.basic /-! # p-adic numbers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the `p`-adic numbers (rationals) `ℚ_[p]` as the completion of `ℚ` with respect to the `p`-adic norm. We show that the `p`-adic norm on `ℚ` extends to `ℚ_[p]`, that `ℚ` is embedded in `ℚ_[p]`, diff --git a/src/order/interval.lean b/src/order/interval.lean index ece0f3aee0dea..bdaa1250c6a6d 100644 --- a/src/order/interval.lean +++ b/src/order/interval.lean @@ -10,6 +10,9 @@ import data.set_like.basic /-! # Order intervals +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines (nonempty) closed intervals in an order (see `set.Icc`). This is a prototype for interval arithmetic. diff --git a/src/ring_theory/algebraic.lean b/src/ring_theory/algebraic.lean index 603178ee4de13..5c629e47a50cf 100644 --- a/src/ring_theory/algebraic.lean +++ b/src/ring_theory/algebraic.lean @@ -11,6 +11,9 @@ import data.polynomial.integral_normalization /-! # Algebraic elements and algebraic extensions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + An element of an R-algebra is algebraic over R if it is the root of a nonzero polynomial. An R-algebra is algebraic over R if and only if all its elements are algebraic over R. The main result in this file proves transitivity of algebraicity: diff --git a/src/ring_theory/artinian.lean b/src/ring_theory/artinian.lean index ed598587452c9..c64f5903be576 100644 --- a/src/ring_theory/artinian.lean +++ b/src/ring_theory/artinian.lean @@ -9,6 +9,9 @@ import data.set_like.fintype /-! # Artinian rings and modules +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A module satisfying these equivalent conditions is said to be an *Artinian* R-module if every decreasing chain of submodules is eventually constant, or equivalently, diff --git a/src/ring_theory/discrete_valuation_ring/basic.lean b/src/ring_theory/discrete_valuation_ring/basic.lean index 5c2df165d347a..27a78c164a18b 100644 --- a/src/ring_theory/discrete_valuation_ring/basic.lean +++ b/src/ring_theory/discrete_valuation_ring/basic.lean @@ -13,6 +13,9 @@ import linear_algebra.adic_completion /-! # Discrete valuation rings +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines discrete valuation rings (DVRs) and develops a basic interface for them. diff --git a/src/ring_theory/graded_algebra/radical.lean b/src/ring_theory/graded_algebra/radical.lean index a4953692a153b..9d6ebfc071937 100644 --- a/src/ring_theory/graded_algebra/radical.lean +++ b/src/ring_theory/graded_algebra/radical.lean @@ -12,6 +12,9 @@ This file contains a proof that the radical of any homogeneous ideal is a homoge ## Main statements +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + * `ideal.is_homogeneous.is_prime_iff`: for any `I : ideal A`, if `I` is homogeneous, then `I` is prime if and only if `I` is homogeneously prime, i.e. `I ≠ ⊤` and if `x, y` are homogeneous elements such that `x * y ∈ I`, then at least one of `x,y` is in `I`. diff --git a/src/ring_theory/integral_closure.lean b/src/ring_theory/integral_closure.lean index aadaeeac3f8bf..fa887f467ad12 100644 --- a/src/ring_theory/integral_closure.lean +++ b/src/ring_theory/integral_closure.lean @@ -15,6 +15,9 @@ import ring_theory.tensor_product /-! # Integral closure of a subring. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + If A is an R-algebra then `a : A` is integral over R if it is a root of a monic polynomial with coefficients in R. Enough theory is developed to prove that integral elements form a sub-R-algebra of A. diff --git a/src/ring_theory/localization/as_subring.lean b/src/ring_theory/localization/as_subring.lean index a233c440442ca..516c28af1768c 100644 --- a/src/ring_theory/localization/as_subring.lean +++ b/src/ring_theory/localization/as_subring.lean @@ -9,6 +9,9 @@ import ring_theory.localization.localization_localization # Localizations of domains as subalgebras of the fraction field. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given a domain `A` with fraction field `K`, and a submonoid `S` of `A` which does not contain zero, this file constructs the localization of `A` at `S` as a subalgebra of the field `K` over `A`. diff --git a/src/ring_theory/localization/away/basic.lean b/src/ring_theory/localization/away/basic.lean index bcdfee9da33f6..417bfe02d192c 100644 --- a/src/ring_theory/localization/away/basic.lean +++ b/src/ring_theory/localization/away/basic.lean @@ -9,6 +9,9 @@ import ring_theory.localization.basic /-! # Localizations away from an element +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions * `is_localization.away (x : R) S` expresses that `S` is a localization away from `x`, as an diff --git a/src/ring_theory/localization/cardinality.lean b/src/ring_theory/localization/cardinality.lean index 84dc211527724..4451cea455ffb 100644 --- a/src/ring_theory/localization/cardinality.lean +++ b/src/ring_theory/localization/cardinality.lean @@ -9,6 +9,9 @@ import ring_theory.artinian /-! # Cardinality of localizations +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we establish the cardinality of localizations. In most cases, a localization has cardinality equal to the base ring. If there are zero-divisors, however, this is no longer true - for example, `zmod 6` localized at `{2, 4}` is equal to `zmod 3`, and if you have zero in your diff --git a/src/ring_theory/localization/localization_localization.lean b/src/ring_theory/localization/localization_localization.lean index 86ed623a3cb6c..5cb930a24a433 100644 --- a/src/ring_theory/localization/localization_localization.lean +++ b/src/ring_theory/localization/localization_localization.lean @@ -10,6 +10,9 @@ import ring_theory.localization.fraction_ring /-! # Localizations of localizations +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Implementation notes See `src/ring_theory/localization/basic.lean` for a design overview. diff --git a/src/ring_theory/power_series/basic.lean b/src/ring_theory/power_series/basic.lean index 591e43568d3f0..78e1c3e211180 100644 --- a/src/ring_theory/power_series/basic.lean +++ b/src/ring_theory/power_series/basic.lean @@ -15,6 +15,9 @@ import tactic.linarith /-! # Formal power series +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines (multivariate) formal power series and develops the basic properties of these objects. diff --git a/src/ring_theory/power_series/well_known.lean b/src/ring_theory/power_series/well_known.lean index 4d6decaf50f78..95b14d16c5a85 100644 --- a/src/ring_theory/power_series/well_known.lean +++ b/src/ring_theory/power_series/well_known.lean @@ -10,6 +10,9 @@ import algebra.big_operators.nat_antidiagonal /-! # Definition of well-known power series +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define the following power series: * `power_series.inv_units_sub`: given `u : Rˣ`, this is the series for `1 / (u - x)`. diff --git a/src/ring_theory/witt_vector/witt_polynomial.lean b/src/ring_theory/witt_vector/witt_polynomial.lean index 59523dec970ed..9adc726f54468 100644 --- a/src/ring_theory/witt_vector/witt_polynomial.lean +++ b/src/ring_theory/witt_vector/witt_polynomial.lean @@ -14,6 +14,9 @@ import data.zmod.basic /-! # Witt polynomials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + To endow `witt_vector p R` with a ring structure, we need to study the so-called Witt polynomials. diff --git a/src/topology/metric_space/metrizable.lean b/src/topology/metric_space/metrizable.lean index 7fb64fc6a9b69..8a016f62e60c7 100644 --- a/src/topology/metric_space/metrizable.lean +++ b/src/topology/metric_space/metrizable.lean @@ -11,6 +11,9 @@ import topology.uniform_space.cauchy /-! # Metrizability of a T₃ topological space with second countable topology +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define metrizable topological spaces, i.e., topological spaces for which there exists a metric space structure that generates the same topology. diff --git a/src/topology/metric_space/thickened_indicator.lean b/src/topology/metric_space/thickened_indicator.lean index 71d0829959d5b..001701e9833f6 100644 --- a/src/topology/metric_space/thickened_indicator.lean +++ b/src/topology/metric_space/thickened_indicator.lean @@ -10,6 +10,9 @@ import topology.metric_space.hausdorff_distance /-! # Thickened indicators +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + 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 diff --git a/src/topology/sheaves/punit.lean b/src/topology/sheaves/punit.lean index c036317c91fdf..67081ab48be0e 100644 --- a/src/topology/sheaves/punit.lean +++ b/src/topology/sheaves/punit.lean @@ -8,6 +8,9 @@ import topology.sheaves.sheaf_condition.sites /-! # Presheaves on punit +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Presheaves on punit satisfy sheaf condition iff its value at empty set is a terminal object. -/ diff --git a/src/topology/urysohns_bounded.lean b/src/topology/urysohns_bounded.lean index 1d497be696a06..c0ae5baa0a4f5 100644 --- a/src/topology/urysohns_bounded.lean +++ b/src/topology/urysohns_bounded.lean @@ -9,6 +9,9 @@ import topology.continuous_function.bounded /-! # Urysohn's lemma for bounded continuous functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we reformulate Urysohn's lemma `exists_continuous_zero_one_of_closed` in terms of bounded continuous functions `X →ᵇ ℝ`. These lemmas live in a separate file because `topology.continuous_function.bounded` imports too many other files. diff --git a/src/topology/vector_bundle/basic.lean b/src/topology/vector_bundle/basic.lean index 67da394c060e5..483fa7312e6ba 100644 --- a/src/topology/vector_bundle/basic.lean +++ b/src/topology/vector_bundle/basic.lean @@ -10,6 +10,9 @@ import topology.fiber_bundle.basic /-! # Vector bundles +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define (topological) vector bundles. Let `B` be the base space, let `F` be a normed space over a normed field `R`, and let diff --git a/src/topology/vector_bundle/constructions.lean b/src/topology/vector_bundle/constructions.lean index f03647810d1f8..08067b91091e8 100644 --- a/src/topology/vector_bundle/constructions.lean +++ b/src/topology/vector_bundle/constructions.lean @@ -9,6 +9,9 @@ import topology.vector_bundle.basic /-! # Standard constructions on vector bundles +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains several standard constructions on vector bundles: * `bundle.trivial.vector_bundle 𝕜 B F`: the trivial vector bundle with scalar field `𝕜` and model From cb9077f70849a93b19e01903d497029014d76e35 Mon Sep 17 00:00:00 2001 From: leanprover-community-bot Date: Tue, 23 May 2023 04:16:26 +0000 Subject: [PATCH 1058/1291] chore(scripts): update nolints.txt (#19065) I am happy to remove some nolints for you! --- scripts/nolints.txt | 20 -------------------- scripts/style-exceptions.txt | 4 ++-- 2 files changed, 2 insertions(+), 22 deletions(-) diff --git a/scripts/nolints.txt b/scripts/nolints.txt index a2a306eb8dfd1..c9d8104dc89b1 100644 --- a/scripts/nolints.txt +++ b/scripts/nolints.txt @@ -179,10 +179,6 @@ apply_nolint list.sublists_aux doc_blame apply_nolint list.sublists_aux₁ doc_blame apply_nolint list.traverse doc_blame --- data/matrix/basic.lean -apply_nolint matrix.mul_submatrix_one fintype_finite -apply_nolint matrix.one_submatrix_mul fintype_finite - -- data/matrix/basis.lean apply_nolint matrix.induction_on fintype_finite apply_nolint matrix.induction_on' fintype_finite @@ -314,22 +310,6 @@ apply_nolint computation.parallel.aux1 doc_blame apply_nolint computation.parallel.aux2 doc_blame apply_nolint computation.parallel_rec doc_blame --- data/seq/seq.lean -apply_nolint seq.bisim_o doc_blame -apply_nolint seq.corec.F doc_blame -apply_nolint seq.is_bisimulation doc_blame -apply_nolint seq.mem doc_blame - --- data/seq/wseq.lean -apply_nolint wseq.bisim_o doc_blame -apply_nolint wseq.destruct_append.aux doc_blame -apply_nolint wseq.destruct_join.aux doc_blame -apply_nolint wseq.drop.aux doc_blame -apply_nolint wseq.lift_rel_o doc_blame -apply_nolint wseq.mem doc_blame -apply_nolint wseq.tail.aux doc_blame -apply_nolint wseq.think_congr unused_arguments - -- data/stream/defs.lean apply_nolint stream.corec doc_blame apply_nolint stream.corec' doc_blame diff --git a/scripts/style-exceptions.txt b/scripts/style-exceptions.txt index 3690c941e9416..c8f9e355af435 100644 --- a/scripts/style-exceptions.txt +++ b/scripts/style-exceptions.txt @@ -4,9 +4,9 @@ src/control/monad/writer.lean : line 11 : ERR_MOD : Module docstring missing, or src/control/traversable/derive.lean : line 11 : ERR_MOD : Module docstring missing, or too late src/data/array/lemmas.lean : line 9 : ERR_MOD : Module docstring missing, or too late src/data/bitvec/basic.lean : line 11 : ERR_MOD : Module docstring missing, or too late -src/data/buffer/basic.lean : line 11 : ERR_MOD : Module docstring missing, or too late +src/data/buffer/basic.lean : line 12 : ERR_MOD : Module docstring missing, or too late src/data/qpf/multivariate/basic.lean : line 76 : ERR_LIN : Line has more than 100 characters -src/data/qpf/univariate/basic.lean : line 35 : ERR_LIN : Line has more than 100 characters +src/data/qpf/univariate/basic.lean : line 38 : ERR_LIN : Line has more than 100 characters src/data/rbmap/basic.lean : line 8 : ERR_MOD : Module docstring missing, or too late src/data/rbmap/default.lean : line 9 : ERR_MOD : Module docstring missing, or too late src/data/rbtree/basic.lean : line 10 : ERR_MOD : Module docstring missing, or too late From 05b93a585ac481dcb91bebd6d210d60996dc2a60 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 23 May 2023 05:24:35 +0000 Subject: [PATCH 1059/1291] feat: update to Lean 3.51.0 (#19060) This version: * Propagates goal tags through `unfreezingI` * Uses a more recent set of emscripten build tools It's hard to know whether this will break the web editor, other than just pushing and finding out. --- leanpkg.toml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/leanpkg.toml b/leanpkg.toml index 98c2e9e127eb3..06b84b6c77b86 100644 --- a/leanpkg.toml +++ b/leanpkg.toml @@ -1,7 +1,7 @@ [package] name = "mathlib" version = "0.1" -lean_version = "leanprover-community/lean:3.50.3" +lean_version = "leanprover-community/lean:3.51.0" path = "src" [dependencies] From c0d694db494dd4f9aa57f2714b6e4c82b4ebc113 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Tue, 23 May 2023 08:35:45 +0000 Subject: [PATCH 1060/1291] feat(probability/kernel/integral_comp_prod): Bochner integral against the composition-product of two kernels (#18976) --- src/probability/kernel/composition.lean | 130 ++++++++ .../kernel/integral_comp_prod.lean | 280 ++++++++++++++++++ src/probability/kernel/with_density.lean | 11 + 3 files changed, 421 insertions(+) create mode 100644 src/probability/kernel/integral_comp_prod.lean diff --git a/src/probability/kernel/composition.lean b/src/probability/kernel/composition.lean index c77ed1f7c9dda..ced5d02480471 100644 --- a/src/probability/kernel/composition.lean +++ b/src/probability/kernel/composition.lean @@ -223,6 +223,101 @@ lemma comp_prod_apply (κ : kernel α β) [is_s_finite_kernel κ] (η : kernel ( (κ ⊗ₖ η) a s = ∫⁻ b, η (a, b) {c | (b, c) ∈ s} ∂(κ a) := comp_prod_apply_eq_comp_prod_fun κ η a hs +lemma le_comp_prod_apply (κ : kernel α β) [is_s_finite_kernel κ] (η : kernel (α × β) γ) + [is_s_finite_kernel η] (a : α) (s : set (β × γ)) : + ∫⁻ b, η (a, b) {c | (b, c) ∈ s} ∂(κ a) ≤ (κ ⊗ₖ η) a s := +calc ∫⁻ b, η (a, b) {c | (b, c) ∈ s} ∂(κ a) + ≤ ∫⁻ b, η (a, b) {c | (b, c) ∈ (to_measurable ((κ ⊗ₖ η) a) s)} ∂(κ a) : + lintegral_mono (λ b, measure_mono (λ _ h_mem, subset_to_measurable _ _ h_mem)) +... = (κ ⊗ₖ η) a (to_measurable ((κ ⊗ₖ η) a) s) : + (kernel.comp_prod_apply_eq_comp_prod_fun κ η a (measurable_set_to_measurable _ _)).symm +... = (κ ⊗ₖ η) a s : measure_to_measurable s + +section ae +/-! ### `ae` filter of the composition-product -/ + +variables {κ : kernel α β} [is_s_finite_kernel κ] {η : kernel (α × β) γ} [is_s_finite_kernel η] + {a : α} + +lemma ae_kernel_lt_top (a : α) (h2s : (κ ⊗ₖ η) a s ≠ ∞) : + ∀ᵐ b ∂(κ a), η (a, b) (prod.mk b ⁻¹' s) < ∞ := +begin + let t := to_measurable ((κ ⊗ₖ η) a) s, + have : ∀ (b : β), η (a, b) (prod.mk b ⁻¹' s) ≤ η (a, b) (prod.mk b ⁻¹' t), + { exact λ b, measure_mono (set.preimage_mono (subset_to_measurable _ _)), }, + have ht : measurable_set t := measurable_set_to_measurable _ _, + have h2t : (κ ⊗ₖ η) a t ≠ ∞, by rwa measure_to_measurable, + have ht_lt_top : ∀ᵐ b ∂(κ a), η (a, b) (prod.mk b ⁻¹' t) < ∞, + { rw kernel.comp_prod_apply _ _ _ ht at h2t, + exact ae_lt_top (kernel.measurable_kernel_prod_mk_left' ht a) h2t, }, + filter_upwards [ht_lt_top] with b hb, + exact (this b).trans_lt hb, +end + +lemma comp_prod_null (a : α) (hs : measurable_set s) : + (κ ⊗ₖ η) a s = 0 ↔ (λ b, η (a, b) (prod.mk b ⁻¹' s)) =ᵐ[κ a] 0 := +begin + rw [kernel.comp_prod_apply _ _ _ hs, lintegral_eq_zero_iff], + { refl, }, + { exact kernel.measurable_kernel_prod_mk_left' hs a, }, +end + +lemma ae_null_of_comp_prod_null (h : (κ ⊗ₖ η) a s = 0) : + (λ b, η (a, b) (prod.mk b ⁻¹' s)) =ᵐ[κ a] 0 := +begin + obtain ⟨t, hst, mt, ht⟩ := exists_measurable_superset_of_null h, + simp_rw [comp_prod_null a mt] at ht, + rw [filter.eventually_le_antisymm_iff], + exact ⟨filter.eventually_le.trans_eq + (filter.eventually_of_forall $ λ x, (measure_mono (set.preimage_mono hst) : _)) ht, + filter.eventually_of_forall $ λ x, zero_le _⟩ +end + +lemma ae_ae_of_ae_comp_prod {p : β × γ → Prop} (h : ∀ᵐ bc ∂((κ ⊗ₖ η) a), p bc) : + ∀ᵐ b ∂(κ a), ∀ᵐ c ∂(η (a, b)), p (b, c) := +ae_null_of_comp_prod_null h + +end ae + +section restrict + +variables {κ : kernel α β} [is_s_finite_kernel κ] {η : kernel (α × β) γ} [is_s_finite_kernel η] + {a : α} + +lemma comp_prod_restrict {s : set β} {t : set γ} (hs : measurable_set s) (ht : measurable_set t) : + (kernel.restrict κ hs) ⊗ₖ (kernel.restrict η ht) = kernel.restrict (κ ⊗ₖ η) (hs.prod ht) := +begin + ext a u hu : 2, + rw [comp_prod_apply _ _ _ hu, restrict_apply' _ _ _ hu, + comp_prod_apply _ _ _ (hu.inter (hs.prod ht))], + simp only [kernel.restrict_apply, measure.restrict_apply' ht, set.mem_inter_iff, + set.prod_mk_mem_set_prod_eq], + have : ∀ b, η (a, b) {c : γ | (b, c) ∈ u ∧ b ∈ s ∧ c ∈ t} + = s.indicator (λ b, η (a, b) ({c : γ | (b, c) ∈ u} ∩ t)) b, + { intro b, + classical, + rw set.indicator_apply, + split_ifs with h, + { simp only [h, true_and], + refl, }, + { simp only [h, false_and, and_false, set.set_of_false, measure_empty], }, }, + simp_rw this, + rw lintegral_indicator _ hs, +end + +lemma comp_prod_restrict_left {s : set β} (hs : measurable_set s) : + (kernel.restrict κ hs) ⊗ₖ η = kernel.restrict (κ ⊗ₖ η) (hs.prod measurable_set.univ) := +by { rw ← comp_prod_restrict, congr, exact kernel.restrict_univ.symm, } + +lemma comp_prod_restrict_right {t : set γ} (ht : measurable_set t) : + κ ⊗ₖ (kernel.restrict η ht) = kernel.restrict (κ ⊗ₖ η) (measurable_set.univ.prod ht) := +by { rw ← comp_prod_restrict, congr, exact kernel.restrict_univ.symm, } + +end restrict + +section lintegral +/-! ### Lebesgue integral -/ + /-- Lebesgue integral against the composition-product of two kernels. -/ theorem lintegral_comp_prod' (κ : kernel α β) [is_s_finite_kernel κ] (η : kernel (α × β) γ) [is_s_finite_kernel η] (a : α) {f : β → γ → ℝ≥0∞} (hf : measurable (function.uncurry f)) : @@ -290,6 +385,41 @@ begin { simp_rw [g, function.uncurry_curry], exact hf, }, end +/-- Lebesgue integral against the composition-product of two kernels. -/ +lemma lintegral_comp_prod₀ (κ : kernel α β) [is_s_finite_kernel κ] (η : kernel (α × β) γ) + [is_s_finite_kernel η] (a : α) {f : β × γ → ℝ≥0∞} (hf : ae_measurable f ((κ ⊗ₖ η) a)) : + ∫⁻ z, f z ∂((κ ⊗ₖ η) a) = ∫⁻ x, ∫⁻ y, f (x, y) ∂(η (a, x)) ∂(κ a) := +begin + have A : ∫⁻ z, f z ∂((κ ⊗ₖ η) a) = ∫⁻ z, hf.mk f z ∂((κ ⊗ₖ η) a) := + lintegral_congr_ae hf.ae_eq_mk, + have B : ∫⁻ x, ∫⁻ y, f (x, y) ∂(η (a, x)) ∂(κ a) = ∫⁻ x, ∫⁻ y, hf.mk f (x, y) ∂(η (a, x)) ∂(κ a), + { apply lintegral_congr_ae, + filter_upwards [ae_ae_of_ae_comp_prod hf.ae_eq_mk] with _ ha using lintegral_congr_ae ha, }, + rw [A, B, lintegral_comp_prod], + exact hf.measurable_mk, +end + +lemma set_lintegral_comp_prod (κ : kernel α β) [is_s_finite_kernel κ] (η : kernel (α × β) γ) + [is_s_finite_kernel η] (a : α) {f : β × γ → ℝ≥0∞} (hf : measurable f) + {s : set β} {t : set γ} (hs : measurable_set s) (ht : measurable_set t) : + ∫⁻ z in s ×ˢ t, f z ∂((κ ⊗ₖ η) a) = ∫⁻ x in s, ∫⁻ y in t, f (x, y) ∂(η (a, x)) ∂(κ a) := +by simp_rw [← kernel.restrict_apply (κ ⊗ₖ η) (hs.prod ht), ← comp_prod_restrict, + lintegral_comp_prod _ _ _ hf, kernel.restrict_apply] + +lemma set_lintegral_comp_prod_univ_right (κ : kernel α β) [is_s_finite_kernel κ] + (η : kernel (α × β) γ) [is_s_finite_kernel η] (a : α) {f : β × γ → ℝ≥0∞} (hf : measurable f) + {s : set β} (hs : measurable_set s) : + ∫⁻ z in s ×ˢ set.univ, f z ∂((κ ⊗ₖ η) a) = ∫⁻ x in s, ∫⁻ y, f (x, y) ∂(η (a, x)) ∂(κ a) := +by simp_rw [set_lintegral_comp_prod κ η a hf hs measurable_set.univ, measure.restrict_univ] + +lemma set_lintegral_comp_prod_univ_left (κ : kernel α β) [is_s_finite_kernel κ] + (η : kernel (α × β) γ) [is_s_finite_kernel η] (a : α) {f : β × γ → ℝ≥0∞} (hf : measurable f) + {t : set γ} (ht : measurable_set t) : + ∫⁻ z in set.univ ×ˢ t, f z ∂((κ ⊗ₖ η) a) = ∫⁻ x, ∫⁻ y in t, f (x, y) ∂(η (a, x)) ∂(κ a) := +by simp_rw [set_lintegral_comp_prod κ η a hf measurable_set.univ ht, measure.restrict_univ] + +end lintegral + lemma comp_prod_eq_tsum_comp_prod (κ : kernel α β) [is_s_finite_kernel κ] (η : kernel (α × β) γ) [is_s_finite_kernel η] (a : α) (hs : measurable_set s) : (κ ⊗ₖ η) a s = ∑' (n m : ℕ), (seq κ n ⊗ₖ seq η m) a s := diff --git a/src/probability/kernel/integral_comp_prod.lean b/src/probability/kernel/integral_comp_prod.lean new file mode 100644 index 0000000000000..7b6917274b16f --- /dev/null +++ b/src/probability/kernel/integral_comp_prod.lean @@ -0,0 +1,280 @@ +/- +Copyright (c) 2023 Rémy Degenne. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Rémy Degenne +-/ +import probability.kernel.composition +import measure_theory.integral.set_integral + +/-! +# Bochner integral of a function against the composition-product of two kernels + +We prove properties of the composition-product of two kernels. If `κ` is an s-finite kernel from +`α` to `β` and `η` is an s-finite kernel from `α × β` to `γ`, we can form their composition-product +`κ ⊗ₖ η : kernel α (β × γ)`. We proved in `probability.kernel.lintegral_comp_prod` that it verifies +`∫⁻ bc, f bc ∂((κ ⊗ₖ η) a) = ∫⁻ b, ∫⁻ c, f (b, c) ∂(η (a, b)) ∂(κ a)`. In this file, we prove the +same equality for the Bochner integral. + +## Main statements + +* `probability_theory.integral_comp_prod`: the integral against the composition-product is + `∫ z, f z ∂((κ ⊗ₖ η) a) = ∫ x, ∫ y, f (x, y) ∂(η (a, x)) ∂(κ a)` + +## Implementation details + +This file is to a large extent a copy of part of `measure_theory.constructions.prod`. The product of +two measures is a particular case of composition-product of kernels and it turns out that once the +measurablity of the Lebesgue integral of a kernel is proved, almost all proofs about integrals +against products of measures extend with minimal modifications to the composition-product of two +kernels. +-/ + +noncomputable theory +open_locale topology ennreal measure_theory probability_theory +open set function real ennreal measure_theory filter probability_theory probability_theory.kernel + +variables {α β γ E : Type*} + {mα : measurable_space α} {mβ : measurable_space β} {mγ : measurable_space γ} + [normed_add_comm_group E] + {κ : kernel α β} [is_s_finite_kernel κ] + {η : kernel (α × β) γ} [is_s_finite_kernel η] + {a : α} + +namespace probability_theory + +lemma has_finite_integral_prod_mk_left (a : α) {s : set (β × γ)} (h2s : (κ ⊗ₖ η) a s ≠ ∞) : + has_finite_integral (λ b, (η (a, b) (prod.mk b ⁻¹' s)).to_real) (κ a) := +begin + let t := to_measurable ((κ ⊗ₖ η) a) s, + simp_rw [has_finite_integral, ennnorm_eq_of_real to_real_nonneg], + calc ∫⁻ b, ennreal.of_real (η (a, b) (prod.mk b ⁻¹' s)).to_real ∂(κ a) + ≤ ∫⁻ b, η (a, b) (prod.mk b ⁻¹' t) ∂(κ a) : + begin + refine lintegral_mono_ae _, + filter_upwards [ae_kernel_lt_top a h2s] with b hb, + rw of_real_to_real hb.ne, + exact measure_mono (preimage_mono (subset_to_measurable _ _)), + end + ... ≤ (κ ⊗ₖ η) a t : le_comp_prod_apply _ _ _ _ + ... = (κ ⊗ₖ η) a s : measure_to_measurable s + ... < ⊤ : h2s.lt_top, +end + +lemma integrable_kernel_prod_mk_left (a : α) + {s : set (β × γ)} (hs : measurable_set s) (h2s : (κ ⊗ₖ η) a s ≠ ∞) : + integrable (λ b, (η (a, b) (prod.mk b ⁻¹' s)).to_real) (κ a) := +begin + split, + { exact (measurable_kernel_prod_mk_left' hs a).ennreal_to_real.ae_strongly_measurable }, + { exact has_finite_integral_prod_mk_left a h2s, }, +end + +lemma _root_.measure_theory.ae_strongly_measurable.integral_kernel_comp_prod + [normed_space ℝ E] [complete_space E] + ⦃f : β × γ → E⦄ (hf : ae_strongly_measurable f ((κ ⊗ₖ η) a)) : + ae_strongly_measurable (λ x, ∫ y, f (x, y) ∂(η (a, x))) (κ a) := +⟨λ x, ∫ y, hf.mk f (x, y) ∂(η (a, x)), hf.strongly_measurable_mk.integral_kernel_prod_right'', + by { filter_upwards [ae_ae_of_ae_comp_prod hf.ae_eq_mk] with _ hx using integral_congr_ae hx }⟩ + +lemma _root_.measure_theory.ae_strongly_measurable.comp_prod_mk_left + {δ : Type*} [topological_space δ] {f : β × γ → δ} + (hf : ae_strongly_measurable f ((κ ⊗ₖ η) a)) : + ∀ᵐ x ∂(κ a), ae_strongly_measurable (λ y, f (x, y)) (η (a, x)) := +by filter_upwards [ae_ae_of_ae_comp_prod hf.ae_eq_mk] with x hx using + ⟨λ y, hf.mk f (x, y), hf.strongly_measurable_mk.comp_measurable measurable_prod_mk_left, hx⟩ + +/-! ### Integrability -/ + +lemma has_finite_integral_comp_prod_iff ⦃f : β × γ → E⦄ (h1f : strongly_measurable f) : + has_finite_integral f ((κ ⊗ₖ η) a) + ↔ (∀ᵐ x ∂(κ a), has_finite_integral (λ y, f (x, y)) (η (a, x))) ∧ + has_finite_integral (λ x, ∫ y, ‖f (x, y)‖ ∂(η (a, x))) (κ a) := +begin + simp only [has_finite_integral], + rw kernel.lintegral_comp_prod _ _ _ h1f.ennnorm, + have : ∀ x, ∀ᵐ y ∂(η (a, x)), 0 ≤ ‖f (x, y)‖ := λ x, eventually_of_forall (λ y, norm_nonneg _), + simp_rw [integral_eq_lintegral_of_nonneg_ae (this _) + (h1f.norm.comp_measurable measurable_prod_mk_left).ae_strongly_measurable, + ennnorm_eq_of_real to_real_nonneg, of_real_norm_eq_coe_nnnorm], + have : ∀ {p q r : Prop} (h1 : r → p), (r ↔ p ∧ q) ↔ (p → (r ↔ q)) := + λ p q r h1, by rw [← and.congr_right_iff, and_iff_right_of_imp h1], + rw [this], + { intro h2f, rw lintegral_congr_ae, + refine h2f.mp _, apply eventually_of_forall, intros x hx, dsimp only, + rw [of_real_to_real], rw [← lt_top_iff_ne_top], exact hx }, + { intro h2f, refine ae_lt_top _ h2f.ne, exact h1f.ennnorm.lintegral_kernel_prod_right'' }, +end + +lemma has_finite_integral_comp_prod_iff' ⦃f : β × γ → E⦄ + (h1f : ae_strongly_measurable f ((κ ⊗ₖ η) a)) : + has_finite_integral f ((κ ⊗ₖ η) a) + ↔ (∀ᵐ x ∂(κ a), has_finite_integral (λ y, f (x, y)) (η (a, x))) ∧ + has_finite_integral (λ x, ∫ y, ‖f (x, y)‖ ∂(η (a, x))) (κ a) := +begin + rw [has_finite_integral_congr h1f.ae_eq_mk, + has_finite_integral_comp_prod_iff h1f.strongly_measurable_mk], + apply and_congr, + { apply eventually_congr, + filter_upwards [ae_ae_of_ae_comp_prod h1f.ae_eq_mk.symm], + assume x hx, + exact has_finite_integral_congr hx }, + { apply has_finite_integral_congr, + filter_upwards [ae_ae_of_ae_comp_prod h1f.ae_eq_mk.symm] with _ hx + using integral_congr_ae (eventually_eq.fun_comp hx _), }, +end + +lemma integrable_comp_prod_iff ⦃f : β × γ → E⦄ (hf : ae_strongly_measurable f ((κ ⊗ₖ η) a)) : + integrable f ((κ ⊗ₖ η) a) ↔ + (∀ᵐ x ∂(κ a), integrable (λ y, f (x, y)) (η (a, x))) + ∧ integrable (λ x, ∫ y, ‖f (x, y)‖ ∂(η (a, x))) (κ a) := +by simp only [integrable, has_finite_integral_comp_prod_iff' hf, + hf.norm.integral_kernel_comp_prod, hf, hf.comp_prod_mk_left, eventually_and, true_and] + +lemma _root_.measure_theory.integrable.comp_prod_mk_left_ae + ⦃f : β × γ → E⦄ (hf : integrable f ((κ ⊗ₖ η) a)) : + ∀ᵐ x ∂(κ a), integrable (λ y, f (x, y)) (η (a, x)) := +((integrable_comp_prod_iff hf.ae_strongly_measurable).mp hf).1 + +lemma _root_.measure_theory.integrable.integral_norm_comp_prod + ⦃f : β × γ → E⦄ (hf : integrable f ((κ ⊗ₖ η) a)) : + integrable (λ x, ∫ y, ‖f (x, y)‖ ∂(η (a, x))) (κ a) := +((integrable_comp_prod_iff hf.ae_strongly_measurable).mp hf).2 + +lemma _root_.measure_theory.integrable.integral_comp_prod [normed_space ℝ E] [complete_space E] + ⦃f : β × γ → E⦄ (hf : integrable f ((κ ⊗ₖ η) a)) : + integrable (λ x, ∫ y, f (x, y) ∂(η (a, x))) (κ a) := +integrable.mono hf.integral_norm_comp_prod + hf.ae_strongly_measurable.integral_kernel_comp_prod $ + eventually_of_forall $ λ x, (norm_integral_le_integral_norm _).trans_eq $ + (norm_of_nonneg $ integral_nonneg_of_ae $ eventually_of_forall $ + λ y, (norm_nonneg (f (x, y)) : _)).symm + +/-! ### Bochner integral -/ + +variables [normed_space ℝ E] [complete_space E] + {E' : Type*} [normed_add_comm_group E'] [complete_space E'] [normed_space ℝ E'] + +lemma kernel.integral_fn_integral_add ⦃f g : β × γ → E⦄ (F : E → E') + (hf : integrable f ((κ ⊗ₖ η) a)) (hg : integrable g ((κ ⊗ₖ η) a)) : + ∫ x, F (∫ y, f (x, y) + g (x, y) ∂(η (a, x))) ∂(κ a) + = ∫ x, F (∫ y, f (x, y) ∂(η (a, x)) + ∫ y, g (x, y) ∂(η (a, x))) ∂(κ a) := +begin + refine integral_congr_ae _, + filter_upwards [hf.comp_prod_mk_left_ae, hg.comp_prod_mk_left_ae] with _ h2f h2g, + simp [integral_add h2f h2g], +end + +lemma kernel.integral_fn_integral_sub ⦃f g : β × γ → E⦄ (F : E → E') + (hf : integrable f ((κ ⊗ₖ η) a)) (hg : integrable g ((κ ⊗ₖ η) a)) : + ∫ x, F (∫ y, f (x, y) - g (x, y) ∂(η (a, x))) ∂(κ a) + = ∫ x, F (∫ y, f (x, y) ∂(η (a, x)) - ∫ y, g (x, y) ∂(η (a, x))) ∂(κ a) := +begin + refine integral_congr_ae _, + filter_upwards [hf.comp_prod_mk_left_ae, hg.comp_prod_mk_left_ae] with _ h2f h2g, + simp [integral_sub h2f h2g], +end + +lemma kernel.lintegral_fn_integral_sub ⦃f g : β × γ → E⦄ + (F : E → ℝ≥0∞) (hf : integrable f ((κ ⊗ₖ η) a)) (hg : integrable g ((κ ⊗ₖ η) a)) : + ∫⁻ x, F (∫ y, f (x, y) - g (x, y) ∂(η (a, x))) ∂(κ a) + = ∫⁻ x, F (∫ y, f (x, y) ∂(η (a, x)) - ∫ y, g (x, y) ∂(η (a, x))) ∂(κ a) := +begin + refine lintegral_congr_ae _, + filter_upwards [hf.comp_prod_mk_left_ae, hg.comp_prod_mk_left_ae] with _ h2f h2g, + simp [integral_sub h2f h2g], +end + +lemma kernel.integral_integral_add ⦃f g : β × γ → E⦄ + (hf : integrable f ((κ ⊗ₖ η) a)) (hg : integrable g ((κ ⊗ₖ η) a)) : + ∫ x, ∫ y, f (x, y) + g (x, y) ∂(η (a, x)) ∂(κ a) + = ∫ x, ∫ y, f (x, y) ∂(η (a, x)) ∂(κ a) + ∫ x, ∫ y, g (x, y) ∂(η (a, x)) ∂(κ a) := +(kernel.integral_fn_integral_add id hf hg).trans $ + integral_add hf.integral_comp_prod hg.integral_comp_prod + +lemma kernel.integral_integral_add' ⦃f g : β × γ → E⦄ + (hf : integrable f ((κ ⊗ₖ η) a)) (hg : integrable g ((κ ⊗ₖ η) a)) : + ∫ x, ∫ y, (f + g) (x, y) ∂(η (a, x)) ∂(κ a) + = ∫ x, ∫ y, f (x, y) ∂(η (a, x)) ∂(κ a) + ∫ x, ∫ y, g (x, y) ∂(η (a, x)) ∂(κ a) := +kernel.integral_integral_add hf hg + +lemma kernel.integral_integral_sub ⦃f g : β × γ → E⦄ + (hf : integrable f ((κ ⊗ₖ η) a)) (hg : integrable g ((κ ⊗ₖ η) a)) : + ∫ x, ∫ y, f (x, y) - g (x, y) ∂(η (a, x)) ∂(κ a) + = ∫ x, ∫ y, f (x, y) ∂(η (a, x)) ∂(κ a) - ∫ x, ∫ y, g (x, y) ∂(η (a, x)) ∂(κ a) := +(kernel.integral_fn_integral_sub id hf hg).trans $ + integral_sub hf.integral_comp_prod hg.integral_comp_prod + +lemma kernel.integral_integral_sub' ⦃f g : β × γ → E⦄ + (hf : integrable f ((κ ⊗ₖ η) a)) (hg : integrable g ((κ ⊗ₖ η) a)) : + ∫ x, ∫ y, (f - g) (x, y) ∂(η (a, x)) ∂(κ a) + = ∫ x, ∫ y, f (x, y) ∂(η (a, x)) ∂(κ a) - ∫ x, ∫ y, g (x, y) ∂(η (a, x)) ∂(κ a) := +kernel.integral_integral_sub hf hg + +lemma kernel.continuous_integral_integral : + continuous (λ (f : α × β →₁[(κ ⊗ₖ η) a] E), ∫ x, ∫ y, f (x, y) ∂(η (a, x)) ∂(κ a)) := +begin + rw [continuous_iff_continuous_at], intro g, + refine tendsto_integral_of_L1 _ (L1.integrable_coe_fn g).integral_comp_prod + (eventually_of_forall $ λ h, (L1.integrable_coe_fn h).integral_comp_prod) _, + simp_rw [← kernel.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, ‖i (x, y) - g (x, y)‖₊ ∂(η (a, x)) ∂(κ a) }, + swap, { exact λ i, lintegral_mono (λ x, ennnorm_integral_le_lintegral_ennnorm _) }, + show tendsto (λ (i : β × γ →₁[(κ ⊗ₖ η) a] E), + ∫⁻ x, ∫⁻ (y : γ), ‖i (x, y) - g (x, y)‖₊ ∂(η (a, x)) ∂(κ a)) (𝓝 g) (𝓝 0), + have : ∀ (i : α × β →₁[(κ ⊗ₖ η) a] E), measurable (λ z, (‖i z - g z‖₊ : ℝ≥0∞)) := + λ i, ((Lp.strongly_measurable i).sub (Lp.strongly_measurable g)).ennnorm, + simp_rw [← kernel.lintegral_comp_prod _ _ _ (this _), ← L1.of_real_norm_sub_eq_lintegral, + ← of_real_zero], + refine (continuous_of_real.tendsto 0).comp _, + rw [← tendsto_iff_norm_tendsto_zero], + exact tendsto_id +end + +lemma integral_comp_prod : ∀ {f : β × γ → E} (hf : integrable f ((κ ⊗ₖ η) a)), + ∫ z, f z ∂((κ ⊗ₖ η) a) = ∫ x, ∫ y, f (x, y) ∂(η (a, x)) ∂(κ a) := +begin + apply integrable.induction, + { intros c s hs h2s, + simp_rw [integral_indicator hs, ← indicator_comp_right, + function.comp, integral_indicator (measurable_prod_mk_left hs), + measure_theory.set_integral_const, integral_smul_const], + congr' 1, + rw integral_to_real, + rotate, + { exact (kernel.measurable_kernel_prod_mk_left' hs _).ae_measurable, }, + { exact (ae_kernel_lt_top a h2s.ne), }, + rw kernel.comp_prod_apply _ _ _ hs, + refl, }, + { intros f g hfg i_f i_g hf hg, + simp_rw [integral_add' i_f i_g, kernel.integral_integral_add' i_f i_g, hf, hg] }, + { exact is_closed_eq continuous_integral kernel.continuous_integral_integral }, + { intros f g hfg i_f hf, + convert hf using 1, + { exact integral_congr_ae hfg.symm }, + { refine integral_congr_ae _, + refine (ae_ae_of_ae_comp_prod hfg).mp (eventually_of_forall _), + exact λ x hfgx, integral_congr_ae (ae_eq_symm hfgx) } } +end + +lemma set_integral_comp_prod {f : β × γ → E} {s : set β} {t : set γ} + (hs : measurable_set s) (ht : measurable_set t) (hf : integrable_on f (s ×ˢ t) ((κ ⊗ₖ η) a)) : + ∫ z in s ×ˢ t, f z ∂((κ ⊗ₖ η) a) = ∫ x in s, ∫ y in t, f (x, y) ∂(η (a, x)) ∂(κ a) := +begin + rw [← kernel.restrict_apply (κ ⊗ₖ η) (hs.prod ht), ← comp_prod_restrict, integral_comp_prod], + { simp_rw kernel.restrict_apply, }, + { rw [comp_prod_restrict, kernel.restrict_apply], exact hf, }, +end + +lemma set_integral_comp_prod_univ_right (f : β × γ → E) {s : set β} + (hs : measurable_set s) (hf : integrable_on f (s ×ˢ univ) ((κ ⊗ₖ η) a)) : + ∫ z in s ×ˢ univ, f z ∂((κ ⊗ₖ η) a) = ∫ x in s, ∫ y, f (x, y) ∂(η (a, x)) ∂(κ a) := +by simp_rw [set_integral_comp_prod hs measurable_set.univ hf, measure.restrict_univ] + +lemma set_integral_comp_prod_univ_left (f : β × γ → E) {t : set γ} + (ht : measurable_set t) (hf : integrable_on f (univ ×ˢ t) ((κ ⊗ₖ η) a)) : + ∫ z in univ ×ˢ t, f z ∂((κ ⊗ₖ η) a) = ∫ x, ∫ y in t, f (x, y) ∂(η (a, x)) ∂(κ a) := +by simp_rw [set_integral_comp_prod measurable_set.univ ht hf, measure.restrict_univ] + +end probability_theory diff --git a/src/probability/kernel/with_density.lean b/src/probability/kernel/with_density.lean index 5ab8e74f45cfb..2b9b2724412d4 100644 --- a/src/probability/kernel/with_density.lean +++ b/src/probability/kernel/with_density.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ import probability.kernel.measurable_integral +import measure_theory.integral.set_integral /-! # With Density @@ -79,6 +80,16 @@ begin simp_rw pi.mul_apply, end +lemma integral_with_density {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] + [complete_space E] {f : β → E} [is_s_finite_kernel κ] {a : α} + {g : α → β → ℝ≥0} (hg : measurable (function.uncurry g)) : + ∫ b, f b ∂(with_density κ (λ a b, g a b) a) = ∫ b, (g a b) • f b ∂(κ a) := +begin + rw [kernel.with_density_apply, integral_with_density_eq_integral_smul], + { exact measurable.of_uncurry_left hg, }, + { exact measurable_coe_nnreal_ennreal.comp hg, }, +end + lemma with_density_add_left (κ η : kernel α β) [is_s_finite_kernel κ] [is_s_finite_kernel η] (f : α → β → ℝ≥0∞) : with_density (κ + η) f = with_density κ f + with_density η f := From ad84a13c884fd19e286fb7abb36f4b9ba7e2f615 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 23 May 2023 12:33:52 +0000 Subject: [PATCH 1061/1291] feat(analysis/calculus/fderiv/star): derivative of star in star-modules over a trivial star-ring (#19038) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Notably this includes the complex and quaternion conjugates. We need the `has_trivial_star` assumption in order to have `star (c • x) = c • star x` available; which in turn we need because `has_fderiv_at` consumes *linear* maps not *semi-linear* maps. In the absence of an easy way to convert between linear and semi-linear maps, we bundle `star` (again) as a continuous linear equiv as `starL'`. Some alternative approaches are discussed [on Zulip](https://leanprover.zulipchat.com/#narrow/stream/217875-Is-there-code-for-X.3F/topic/Star-semilinear.20maps.20are.20semilinear.20when.20star.20is.20trivial/near/359552581). These API here is just the API for `-` (`neg`), modified by replacing `neg` with `star` and `-f x` with `star (f x)`. Since we require `has_trivial_star` there is no point adding any lemmas for the derivative of `star` on the field itself. --- src/analysis/calculus/deriv.lean | 36 ++++++++++ src/analysis/calculus/fderiv/star.lean | 91 ++++++++++++++++++++++++++ src/topology/algebra/module/star.lean | 16 +++++ 3 files changed, 143 insertions(+) create mode 100644 src/analysis/calculus/fderiv/star.lean diff --git a/src/analysis/calculus/deriv.lean b/src/analysis/calculus/deriv.lean index 178b32111773e..60f5e374fb9d7 100644 --- a/src/analysis/calculus/deriv.lean +++ b/src/analysis/calculus/deriv.lean @@ -7,6 +7,7 @@ import analysis.calculus.fderiv.add import analysis.calculus.fderiv.mul import analysis.calculus.fderiv.equiv import analysis.calculus.fderiv.restrict_scalars +import analysis.calculus.fderiv.star import data.polynomial.algebra_map import data.polynomial.derivative import linear_algebra.affine_space.slope @@ -56,6 +57,7 @@ We also show the existence and compute the derivatives of: - negation - subtraction - multiplication + - star - inverse `x → x⁻¹` - multiplication of two functions in `𝕜 → 𝕜` - multiplication of a function in `𝕜 → 𝕜` and of a function in `𝕜 → E` @@ -1670,6 +1672,40 @@ by simp only [div_eq_mul_inv, deriv_mul_const_field] end division +section star +/-! ### Derivative of `x ↦ star x` -/ + +variables [star_ring 𝕜] [has_trivial_star 𝕜] [star_add_monoid F] [has_continuous_star F] +variable [star_module 𝕜 F] + +protected theorem has_deriv_at_filter.star (h : has_deriv_at_filter f f' x L) : + has_deriv_at_filter (λ x, star (f x)) (star f') x L := +by simpa using h.star.has_deriv_at_filter + +protected theorem has_deriv_within_at.star (h : has_deriv_within_at f f' s x) : + has_deriv_within_at (λ x, star (f x)) (star f') s x := +h.star + +protected theorem has_deriv_at.star (h : has_deriv_at f f' x) : + has_deriv_at (λ x, star (f x)) (star f') x := +h.star + +protected theorem has_strict_deriv_at.star (h : has_strict_deriv_at f f' x) : + has_strict_deriv_at (λ x, star (f x)) (star f') x := +by simpa using h.star.has_strict_deriv_at + +protected lemma deriv_within.star (hxs : unique_diff_within_at 𝕜 s x) : + deriv_within (λ y, star (f y)) s x = star (deriv_within f s x) := +fun_like.congr_fun (fderiv_within_star hxs) _ + +protected lemma deriv.star : deriv (λ y, star (f y)) x = star (deriv f x) := +fun_like.congr_fun fderiv_star _ + +@[simp] protected lemma deriv.star' : deriv (λ y, star (f y)) = (λ x, star (deriv f x)) := +funext $ λ x, deriv.star + +end star + section clm_comp_apply /-! ### Derivative of the pointwise composition/application of continuous linear maps -/ diff --git a/src/analysis/calculus/fderiv/star.lean b/src/analysis/calculus/fderiv/star.lean new file mode 100644 index 0000000000000..556f4934f6562 --- /dev/null +++ b/src/analysis/calculus/fderiv/star.lean @@ -0,0 +1,91 @@ +/- +Copyright (c) 2023 Eric Wieser. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Eric Wieser +-/ +import analysis.calculus.fderiv.linear +import analysis.calculus.fderiv.comp +import analysis.calculus.fderiv.equiv +import analysis.normed_space.star.basic + +/-! +# Star operations on derivatives + +For detailed documentation of the Fréchet derivative, +see the module docstring of `analysis/calculus/fderiv/basic.lean`. + +This file contains the usual formulas (and existence assertions) for the derivative of the star +operation. Note that these only apply when the field that the derivative is respect to has a trivial +star operation; which as should be expected rules out `𝕜 = ℂ`. +-/ + +open_locale classical + + +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] [star_ring 𝕜] [has_trivial_star 𝕜] +variables {E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] +variables {F : Type*} [normed_add_comm_group F] [star_add_monoid F] [normed_space 𝕜 F] + [star_module 𝕜 F] [has_continuous_star F] + +variables {f : E → F} +variables {f' : E →L[𝕜] F} +variables (e : E →L[𝕜] F) +variables {x : E} +variables {s : set E} +variables {L : filter E} + +theorem has_strict_fderiv_at.star (h : has_strict_fderiv_at f f' x) : + has_strict_fderiv_at (λ x, star (f x)) (((starL' 𝕜 : F ≃L[𝕜] F) : F →L[𝕜] F) ∘L f') x := +(starL' 𝕜 : F ≃L[𝕜] F).to_continuous_linear_map.has_strict_fderiv_at.comp x h + +theorem has_fderiv_at_filter.star (h : has_fderiv_at_filter f f' x L) : + has_fderiv_at_filter (λ x, star (f x)) (((starL' 𝕜 : F ≃L[𝕜] F) : F →L[𝕜] F) ∘L f') x L := +(starL' 𝕜 : F ≃L[𝕜] F).to_continuous_linear_map.has_fderiv_at_filter.comp x h filter.tendsto_map + +theorem has_fderiv_within_at.star (h : has_fderiv_within_at f f' s x) : + has_fderiv_within_at (λ x, star (f x)) (((starL' 𝕜 : F ≃L[𝕜] F) : F →L[𝕜] F) ∘L f') s x := +h.star + +theorem has_fderiv_at.star (h : has_fderiv_at f f' x) : + has_fderiv_at (λ x, star (f x)) (((starL' 𝕜 : F ≃L[𝕜] F) : F →L[𝕜] F) ∘L f') x := +h.star + +lemma differentiable_within_at.star (h : differentiable_within_at 𝕜 f s x) : + differentiable_within_at 𝕜 (λ y, star (f y)) s x := +h.has_fderiv_within_at.star.differentiable_within_at + +@[simp] lemma differentiable_within_at_star_iff : + differentiable_within_at 𝕜 (λ y, star (f y)) s x ↔ differentiable_within_at 𝕜 f s x := +(starL' 𝕜 : F ≃L[𝕜] F).comp_differentiable_within_at_iff + +lemma differentiable_at.star (h : differentiable_at 𝕜 f x) : + differentiable_at 𝕜 (λ y, star (f y)) x := +h.has_fderiv_at.star.differentiable_at + +@[simp] lemma differentiable_at_star_iff : + differentiable_at 𝕜 (λ y, star (f y)) x ↔ differentiable_at 𝕜 f x := +(starL' 𝕜 : F ≃L[𝕜] F).comp_differentiable_at_iff + +lemma differentiable_on.star (h : differentiable_on 𝕜 f s) : + differentiable_on 𝕜 (λ y, star (f y)) s := +λ x hx, (h x hx).star + +@[simp] lemma differentiable_on_star_iff : + differentiable_on 𝕜 (λ y, star (f y)) s ↔ differentiable_on 𝕜 f s := +(starL' 𝕜 : F ≃L[𝕜] F).comp_differentiable_on_iff + +lemma differentiable.star (h : differentiable 𝕜 f) : + differentiable 𝕜 (λ y, star (f y)) := +λx, (h x).star + +@[simp] lemma differentiable_star_iff : differentiable 𝕜 (λ y, star (f y)) ↔ differentiable 𝕜 f := +(starL' 𝕜 : F ≃L[𝕜] F).comp_differentiable_iff + +lemma fderiv_within_star (hxs : unique_diff_within_at 𝕜 s x) : + fderiv_within 𝕜 (λ y, star (f y)) s x = + ((starL' 𝕜 : F ≃L[𝕜] F) : F →L[𝕜] F) ∘L fderiv_within 𝕜 f s x := +(starL' 𝕜 : F ≃L[𝕜] F).comp_fderiv_within hxs + +@[simp] lemma fderiv_star : + fderiv 𝕜 (λ y, star (f y)) x = ((starL' 𝕜 : F ≃L[𝕜] F) : F →L[𝕜] F) ∘L fderiv 𝕜 f x := +(starL' 𝕜 : F ≃L[𝕜] F).comp_fderiv diff --git a/src/topology/algebra/module/star.lean b/src/topology/algebra/module/star.lean index 07074a3d670d8..859462c2b60ab 100644 --- a/src/topology/algebra/module/star.lean +++ b/src/topology/algebra/module/star.lean @@ -25,6 +25,22 @@ def starL (R : Type*) {A : Type*} continuous_to_fun := continuous_star, continuous_inv_fun := continuous_star } +-- TODO: this could be replaced with something like `(starL R).restrict_scalarsₛₗ h` if we +-- implemented the idea in +-- https://leanprover.zulipchat.com/#narrow/stream/217875-Is-there-code-for-X.3F/topic/Star-semilinear.20maps.20are.20semilinear.20when.20star.20is.20trivial/near/359557835 +/-- If `A` is a topological module over a commutative `R` with trivial star and compatible actions, +then `star` is a continuous linear equivalence. -/ +@[simps] +def starL' (R : Type*) {A : Type*} + [comm_semiring R] [star_ring R] [has_trivial_star R] [add_comm_monoid A] [star_add_monoid A] + [module R A] [star_module R A] [topological_space A] [has_continuous_star A] : + A ≃L[R] A := +(starL R : A ≃L⋆[R] A).trans + ({ map_smul' := λ r a, by simp [star_ring_end_apply], + continuous_to_fun := continuous_id, + continuous_inv_fun := continuous_id, + ..add_equiv.refl A, } : A ≃L⋆[R] A) + variables (R : Type*) (A : Type*) [semiring R] [star_semigroup R] [has_trivial_star R] [add_comm_group A] [module R A] [star_add_monoid A] [star_module R A] From b602702a58f74f5317862a24893693e80bee6d41 Mon Sep 17 00:00:00 2001 From: Eric Rodriguez Date: Tue, 23 May 2023 14:49:31 +0000 Subject: [PATCH 1062/1291] chore(number_theory/cyclotomic): tidying (#19067) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Now that the splitting field diamond is gone, we can tidy the code around these areas a bit more. This also incidentally removed another diamond (the ℤ-algebra on `cyclotomic_ring`) and so we can remove ugliness related to this. --- src/number_theory/cyclotomic/basic.lean | 8 +++---- .../cyclotomic/primitive_roots.lean | 21 +++++++++---------- 2 files changed, 14 insertions(+), 15 deletions(-) diff --git a/src/number_theory/cyclotomic/basic.lean b/src/number_theory/cyclotomic/basic.lean index 56086e1890cfe..bc009f7e7731f 100644 --- a/src/number_theory/cyclotomic/basic.lean +++ b/src/number_theory/cyclotomic/basic.lean @@ -573,11 +573,11 @@ def cyclotomic_ring : Type w := adjoin A { b : (cyclotomic_field n K) | b ^ (n : namespace cyclotomic_ring -/-- The `A`-algebra structure on `cyclotomic_ring n A K`. -This is not an instance since it causes diamonds when `A = ℤ`. -/ -def algebra_base : algebra A (cyclotomic_ring n A K) := (adjoin A _).algebra +/-- The `A`-algebra structure on `cyclotomic_ring n A K`. -/ +instance algebra_base : algebra A (cyclotomic_ring n A K) := (adjoin A _).algebra -local attribute [instance] cyclotomic_ring.algebra_base +-- Ensure that there is no diamonds with ℤ. +example {n : ℕ+} : cyclotomic_ring.algebra_base n ℤ ℚ = algebra_int _ := rfl instance : no_zero_smul_divisors A (cyclotomic_ring n A K) := (adjoin A _).no_zero_smul_divisors_bot diff --git a/src/number_theory/cyclotomic/primitive_roots.lean b/src/number_theory/cyclotomic/primitive_roots.lean index 802553dac2667..11c48e29db1d8 100644 --- a/src/number_theory/cyclotomic/primitive_roots.lean +++ b/src/number_theory/cyclotomic/primitive_roots.lean @@ -13,14 +13,14 @@ import ring_theory.norm /-! # Primitive roots in cyclotomic fields -If `is_cyclotomic_extension {n} A B`, we define an element `zeta n A B : B` that is (under certain -assumptions) a primitive `n`-root of unity in `B` and we study its properties. We also prove related -theorems under the more general assumption of just being a primitive root, for reasons described -in the implementation details section. +If `is_cyclotomic_extension {n} A B`, we define an element `zeta n A B : B` that is a primitive +`n`th-root of unity in `B` and we study its properties. We also prove related theorems under the +more general assumption of just being a primitive root, for reasons described in the implementation +details section. ## Main definitions * `is_cyclotomic_extension.zeta n A B`: if `is_cyclotomic_extension {n} A B`, than `zeta n A B` - is an element of `B` that plays the role of a primitive `n`-th root of unity. + is a primitive `n`-th root of unity in `B`. * `is_primitive_root.power_basis`: if `K` and `L` are fields such that `is_cyclotomic_extension {n} K L`, then `is_primitive_root.power_basis` gives a K-power basis for L given a primitive root `ζ`. @@ -35,7 +35,7 @@ in the implementation details section. the norm of a primitive root is `1` if `n ≠ 2`. * `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`. + primitive root `ζ`. 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 @@ -47,9 +47,9 @@ in the implementation details section. and `primitive_roots n A` given by the choice of `ζ`. ## Implementation details -`zeta n A B` is defined as any primitive root of unity in `B`, that exists by definition. It is not -true in general that it is a root of `cyclotomic n B`, but this holds if `is_domain B` and -`ne_zero (↑n : B)`. +`zeta n A B` is defined as any primitive root of unity in `B`, - this must exist, by definition of +`is_cyclotomic_extension`. It is not true in general that it is a root of `cyclotomic n B`, +but this holds if `is_domain B` and `ne_zero (↑n : B)`. `zeta n A B` is defined using `exists.some`, which means we cannot control it. For example, in normal mathematics, we can demand that `(zeta p ℤ ℤ[ζₚ] : ℚ(ζₚ))` is equal to @@ -308,8 +308,7 @@ begin simp end -local attribute [instance] is_cyclotomic_extension.finite_dimensional -local attribute [instance] is_cyclotomic_extension.is_galois +open_locale cyclotomic /-- 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 From ba5ff5ad5d120fb0ef094ad2994967e9bfaf5112 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 23 May 2023 16:54:20 +0000 Subject: [PATCH 1063/1291] refactor(analysis/normed_space/basic): generalize some results to actions by normed_rings (#19053) This only moves the very basic lemmas for now. This should be very easy to forward-port: * Let someone port the new file via the normal mechanism * Have them delete the duplicate lemmas that appear in CI A few downstream proofs need some small help with unification, as while the old `normed_space` argument was found by unification, the new `has_bounded_smul` has to be found by typeclass search. --- src/analysis/normed/mul_action.lean | 98 +++++++++++++++++++ src/analysis/normed_space/basic.lean | 56 ++--------- .../normed_space/star/multiplier.lean | 2 +- src/analysis/special_functions/exp.lean | 4 +- .../continuous_function/zero_at_infty.lean | 2 +- 5 files changed, 109 insertions(+), 53 deletions(-) create mode 100644 src/analysis/normed/mul_action.lean diff --git a/src/analysis/normed/mul_action.lean b/src/analysis/normed/mul_action.lean new file mode 100644 index 0000000000000..91e0cc442be98 --- /dev/null +++ b/src/analysis/normed/mul_action.lean @@ -0,0 +1,98 @@ +/- +Copyright (c) 2023 Eric Wieser. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Eric Wieser +-/ +import topology.metric_space.algebra +import analysis.normed.field.basic + +/-! +# Lemmas for `has_bounded_smul` over normed additive groups + +Lemmas which hold only in `normed_space α β` are provided in another file. + +Notably we prove that `non_unital_semi_normed_ring`s have bounded actions by left- and right- +multiplication. This allows downstream files to write general results about `bounded_smul`, and then +deduce `const_mul` and `mul_const` results as an immediate corollary. +-/ + +variables {α β : Type*} + +section seminormed_add_group +variables [seminormed_add_group α] [seminormed_add_group β] [smul_zero_class α β] +variables [has_bounded_smul α β] + +lemma norm_smul_le (r : α) (x : β) : ‖r • x‖ ≤ ‖r‖ * ‖x‖ := +by simpa [smul_zero] using dist_smul_pair r 0 x + +lemma nnnorm_smul_le (r : α) (x : β) : ‖r • x‖₊ ≤ ‖r‖₊ * ‖x‖₊ := +norm_smul_le _ _ + +lemma dist_smul_le (s : α) (x y : β) : dist (s • x) (s • y) ≤ ‖s‖ * dist x y := +by simpa only [dist_eq_norm, sub_zero] using dist_smul_pair s x y + +lemma nndist_smul_le (s : α) (x y : β) : nndist (s • x) (s • y) ≤ ‖s‖₊ * nndist x y := +dist_smul_le s x y + +lemma lipschitz_with_smul (s : α) : lipschitz_with ‖s‖₊ ((•) s : β → β) := +lipschitz_with_iff_dist_le_mul.2 $ dist_smul_le _ + +end seminormed_add_group + +/-- Left multiplication is bounded. -/ +instance non_unital_semi_normed_ring.to_has_bounded_smul [non_unital_semi_normed_ring α] : + has_bounded_smul α α := +{ dist_smul_pair' := λ x y₁ y₂, by simpa [mul_sub, dist_eq_norm] using norm_mul_le x (y₁ - y₂), + dist_pair_smul' := λ x₁ x₂ y, by simpa [sub_mul, dist_eq_norm] using norm_mul_le (x₁ - x₂) y, } + +/-- Right multiplication is bounded. -/ +instance non_unital_semi_normed_ring.to_has_bounded_op_smul [non_unital_semi_normed_ring α] : + has_bounded_smul αᵐᵒᵖ α := +{ dist_smul_pair' := λ x y₁ y₂, + by simpa [sub_mul, dist_eq_norm, mul_comm] using norm_mul_le (y₁ - y₂) x.unop, + dist_pair_smul' := λ x₁ x₂ y, + by simpa [mul_sub, dist_eq_norm, mul_comm] using norm_mul_le y (x₁ - x₂).unop, } + +section semi_normed_ring +variables [semi_normed_ring α] [seminormed_add_comm_group β] [module α β] + +lemma has_bounded_smul.of_norm_smul_le (h : ∀ (r : α) (x : β), ‖r • x‖ ≤ ‖r‖ * ‖x‖) : + has_bounded_smul α β := +{ dist_smul_pair' := λ a b₁ b₂, by simpa [smul_sub, dist_eq_norm] using h a (b₁ - b₂), + dist_pair_smul' := λ a₁ a₂ b, by simpa [sub_smul, dist_eq_norm] using h (a₁ - a₂) b } + +end semi_normed_ring + +section normed_division_ring +variables [normed_division_ring α] [seminormed_add_group β] +variables [mul_action_with_zero α β] [has_bounded_smul α β] + +lemma norm_smul (r : α) (x : β) : ‖r • x‖ = ‖r‖ * ‖x‖ := +begin + by_cases h : r = 0, + { simp [h, zero_smul α x] }, + { refine le_antisymm (norm_smul_le r x) _, + calc ‖r‖ * ‖x‖ = ‖r‖ * ‖r⁻¹ • r • x‖ : by rw [inv_smul_smul₀ h] + ... ≤ ‖r‖ * (‖r⁻¹‖ * ‖r • x‖) : + mul_le_mul_of_nonneg_left (norm_smul_le _ _) (norm_nonneg _) + ... = ‖r • x‖ : + by rw [norm_inv, ← mul_assoc, mul_inv_cancel (mt norm_eq_zero.1 h), one_mul] } +end + +lemma nnnorm_smul (r : α) (x : β) : ‖r • x‖₊ = ‖r‖₊ * ‖x‖₊ := +nnreal.eq $ norm_smul r x + +end normed_division_ring + +section normed_division_ring_module +variables [normed_division_ring α] [seminormed_add_comm_group β] +variables [module α β] [has_bounded_smul α β] + +lemma dist_smul₀ (s : α) (x y : β) : dist (s • x) (s • y) = ‖s‖ * dist x y := +by simp_rw [dist_eq_norm, (norm_smul _ _).symm, smul_sub] + +lemma nndist_smul₀ (s : α) (x y : β) : + nndist (s • x) (s • y) = ‖s‖₊ * nndist x y := +nnreal.eq $ dist_smul₀ s x y + +end normed_division_ring_module diff --git a/src/analysis/normed_space/basic.lean b/src/analysis/normed_space/basic.lean index 3e70aa5b481c3..b75f459755dcf 100644 --- a/src/analysis/normed_space/basic.lean +++ b/src/analysis/normed_space/basic.lean @@ -6,6 +6,7 @@ Authors: Patrick Massot, Johannes Hölzl import algebra.algebra.pi import algebra.algebra.restrict_scalars import analysis.normed.field.basic +import analysis.normed.mul_action import data.real.sqrt import topology.algebra.module.basic @@ -45,31 +46,9 @@ end prio variables [normed_field α] [seminormed_add_comm_group β] --- note: while these are currently strictly weaker than the versions without `le`, they will cease --- to be if we eventually generalize `normed_space` from `normed_field α` to `normed_ring α`. -section le - -lemma norm_smul_le [normed_space α β] (r : α) (x : β) : ‖r • x‖ ≤ ‖r‖ * ‖x‖ := -normed_space.norm_smul_le _ _ - -lemma nnnorm_smul_le [normed_space α β] (s : α) (x : β) : ‖s • x‖₊ ≤ ‖s‖₊ * ‖x‖₊ := -norm_smul_le s x - -lemma dist_smul_le [normed_space α β] (s : α) (x y : β) : dist (s • x) (s • y) ≤ ‖s‖ * dist x y := -by simpa only [dist_eq_norm, ←smul_sub] using norm_smul_le _ _ - -lemma nndist_smul_le [normed_space α β] (s : α) (x y : β) : - nndist (s • x) (s • y) ≤ ‖s‖₊ * nndist x y := -dist_smul_le s x y - -end le - @[priority 100] -- see Note [lower instance priority] instance normed_space.has_bounded_smul [normed_space α β] : has_bounded_smul α β := -{ dist_smul_pair' := λ x y₁ y₂, - by simpa [dist_eq_norm, smul_sub] using norm_smul_le x (y₁ - y₂), - dist_pair_smul' := λ x₁ x₂ y, - by simpa [dist_eq_norm, sub_smul] using norm_smul_le (x₁ - x₂) y } +has_bounded_smul.of_norm_smul_le normed_space.norm_smul_le -- Shortcut instance, as otherwise this will be found by `normed_space.to_module` and be -- noncomputable. @@ -78,17 +57,9 @@ instance : module ℝ ℝ := by apply_instance instance normed_field.to_normed_space : normed_space α α := { norm_smul_le := λ a b, norm_mul_le a b } -lemma norm_smul [normed_space α β] (s : α) (x : β) : ‖s • x‖ = ‖s‖ * ‖x‖ := -begin - by_cases h : s = 0, - { simp [h] }, - { refine le_antisymm (norm_smul_le s x) _, - calc ‖s‖ * ‖x‖ = ‖s‖ * ‖s⁻¹ • s • x‖ : by rw [inv_smul_smul₀ h] - ... ≤ ‖s‖ * (‖s⁻¹‖ * ‖s • x‖) : - mul_le_mul_of_nonneg_left (norm_smul_le _ _) (norm_nonneg _) - ... = ‖s • x‖ : - by rw [norm_inv, ← mul_assoc, mul_inv_cancel (mt norm_eq_zero.1 h), one_mul] } -end +-- shortcut instance +instance normed_field.to_has_bounded_smul : has_bounded_smul α α := +normed_space.has_bounded_smul lemma norm_zsmul (α) [normed_field α] [normed_space α β] (n : ℤ) (x : β) : ‖n • x‖ = ‖(n : α)‖ * ‖x‖ := @@ -102,19 +73,6 @@ lemma inv_norm_smul_mem_closed_unit_ball [normed_space ℝ β] (x : β) : by simp only [mem_closed_ball_zero_iff, norm_smul, norm_inv, norm_norm, ← div_eq_inv_mul, div_self_le_one] -lemma dist_smul₀ [normed_space α β] (s : α) (x y : β) : dist (s • x) (s • y) = ‖s‖ * dist x y := -by simp only [dist_eq_norm, (norm_smul _ _).symm, smul_sub] - -lemma nnnorm_smul [normed_space α β] (s : α) (x : β) : ‖s • x‖₊ = ‖s‖₊ * ‖x‖₊ := -nnreal.eq $ norm_smul s x - -lemma nndist_smul₀ [normed_space α β] (s : α) (x y : β) : - nndist (s • x) (s • y) = ‖s‖₊ * nndist x y := -nnreal.eq $ dist_smul₀ s x y - -lemma lipschitz_with_smul [normed_space α β] (s : α) : lipschitz_with ‖s‖₊ ((•) s : β → β) := -lipschitz_with_iff_dist_le_mul.2 $ λ x y, by rw [dist_smul₀, coe_nnnorm] - lemma norm_smul_of_nonneg [normed_space ℝ β] {t : ℝ} (ht : 0 ≤ t) (x : β) : ‖t • x‖ = t * ‖x‖ := by rw [norm_smul, real.norm_eq_abs, abs_of_nonneg ht] @@ -279,7 +237,7 @@ instance pi.normed_space {E : ι → Type*} [fintype ι] [∀i, seminormed_add_c end } instance mul_opposite.normed_space : normed_space α Eᵐᵒᵖ := -{ norm_smul_le := λ s x, norm_smul_le s x.unop, +{ norm_smul_le := λ s x, (norm_smul_le s x.unop : _), ..mul_opposite.normed_add_comm_group, ..mul_opposite.module _ } @@ -288,7 +246,7 @@ instance submodule.normed_space {𝕜 R : Type*} [has_smul 𝕜 R] [normed_field {E : Type*} [seminormed_add_comm_group E] [normed_space 𝕜 E] [module R E] [is_scalar_tower 𝕜 R E] (s : submodule R E) : normed_space 𝕜 s := -{ norm_smul_le := λc x, norm_smul_le c (x : E) } +{ norm_smul_le := λc x, (norm_smul_le c (x : E) : _) } /-- If there is a scalar `c` with `‖c‖>1`, then any element with nonzero norm can be moved by scalar multiplication to any shell of width `‖c‖`. Also recap information on the norm of diff --git a/src/analysis/normed_space/star/multiplier.lean b/src/analysis/normed_space/star/multiplier.lean index 26302dcbd1276..16a4dd5d64409 100644 --- a/src/analysis/normed_space/star/multiplier.lean +++ b/src/analysis/normed_space/star/multiplier.lean @@ -383,7 +383,7 @@ lemma norm_def' (a : 𝓜(𝕜, A)) : ‖a‖ = ‖a.to_prod_mul_opposite_hom‖ lemma nnnorm_def' (a : 𝓜(𝕜, A)) : ‖a‖₊ = ‖a.to_prod_mul_opposite_hom‖₊ := rfl instance : normed_space 𝕜 𝓜(𝕜, A) := -{ norm_smul_le := λ k a, norm_smul_le k a.to_prod_mul_opposite, +{ norm_smul_le := λ k a, (norm_smul_le k a.to_prod_mul_opposite : _), .. double_centralizer.module } instance : normed_algebra 𝕜 𝓜(𝕜, A) := diff --git a/src/analysis/special_functions/exp.lean b/src/analysis/special_functions/exp.lean index 02e03f130d6b6..24a73fee4cb2a 100644 --- a/src/analysis/special_functions/exp.lean +++ b/src/analysis/special_functions/exp.lean @@ -48,8 +48,8 @@ begin have h_sq : ∀ z, ‖z‖ ≤ 1 → ‖exp (x + z) - exp x‖ ≤ ‖z‖ * ‖exp x‖ + ‖exp x‖ * ‖z‖ ^ 2, { intros z hz, have : ‖exp (x + z) - exp x - z • exp x‖ ≤ ‖exp x‖ * ‖z‖ ^ 2, from exp_bound_sq x z hz, - rw [← sub_le_iff_le_add', ← norm_smul z], - exact (norm_sub_norm_le _ _).trans this, }, + rw [← sub_le_iff_le_add', ← norm_smul z (_ : ℂ)], + exact (norm_sub_norm_le _ _).trans this }, calc ‖exp y - exp x‖ = ‖exp (x + (y - x)) - exp x‖ : by nth_rewrite 0 hy_eq ... ≤ ‖y - x‖ * ‖exp x‖ + ‖exp x‖ * ‖y - x‖ ^ 2 : h_sq (y - x) (hyx.le.trans hr_le) ... ≤ ‖y - x‖ * ‖exp x‖ + ‖exp x‖ * (r * ‖y - x‖) : diff --git a/src/topology/continuous_function/zero_at_infty.lean b/src/topology/continuous_function/zero_at_infty.lean index 2eca0c8ff9a92..f85b63491ed92 100644 --- a/src/topology/continuous_function/zero_at_infty.lean +++ b/src/topology/continuous_function/zero_at_infty.lean @@ -411,7 +411,7 @@ normed_add_comm_group.induced C₀(α, β) (α →ᵇ β) (⟨to_bcf, rfl, λ x lemma norm_to_bcf_eq_norm {f : C₀(α, β)} : ‖f.to_bcf‖ = ‖f‖ := rfl instance : normed_space 𝕜 C₀(α, β) := -{ norm_smul_le := λ k f, norm_smul_le k f.to_bcf } +{ norm_smul_le := λ k f, (norm_smul_le k f.to_bcf : _) } end normed_space From 8d33f09cd7089ecf074b4791907588245aec5d1b Mon Sep 17 00:00:00 2001 From: "Filippo A. E. Nuccio" Date: Tue, 23 May 2023 19:44:12 +0000 Subject: [PATCH 1064/1291] feat(topology/homotopy/H_spaces): define H spaces (#16029) introduce H-spaces and prove basic properties, in particular that every topological group is a H-space and that the path space at a point is a H-space. Co-authored-by: faenuccio Co-authored-by: Filippo A. E. Nuccio Co-authored-by: Junyan Xu --- docs/references.bib | 19 ++- src/topology/homotopy/H_spaces.lean | 211 ++++++++++++++++++++++++++++ 2 files changed, 228 insertions(+), 2 deletions(-) create mode 100644 src/topology/homotopy/H_spaces.lean diff --git a/docs/references.bib b/docs/references.bib index 470569c9a9037..d5c6b40fab9e3 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -1997,6 +1997,22 @@ @Book{ seligman1967 mrreviewer = {R. E. Block} } +@Article{ serre1951, + author = {Serre, Jean-Pierre}, + title = {Homologie singuli\`ere des espaces fibr\'{e}s. + {A}pplications}, + journal = {Ann. of Math. (2)}, + year = {1951}, + volume = {54}, + pages = {425--505}, + issn = {0003-486X}, + doi = {10.2307/1969485}, + fjournal = {Annals of Mathematics. Second Series}, + mrclass = {56.0X}, + mrnumber = {0045386}, + mrreviewer = {W. S. Massey} +} + @Book{ serre1965, author = {Serre, Jean-Pierre}, title = {Complex semisimple {L}ie algebras}, @@ -2007,8 +2023,7 @@ @Book{ serre1965 isbn = {0-387-96569-6}, mrclass = {17-01 (17B20)}, mrnumber = {914496}, - doi = {10.1007/978-1-4757-3910-7}, - url = {https://doi.org/10.1007/978-1-4757-3910-7} + doi = {10.1007/978-1-4757-3910-7} } @Book{ silverman2009, diff --git a/src/topology/homotopy/H_spaces.lean b/src/topology/homotopy/H_spaces.lean new file mode 100644 index 0000000000000..d39b2b8a4b8c0 --- /dev/null +++ b/src/topology/homotopy/H_spaces.lean @@ -0,0 +1,211 @@ +/- +Copyright (c) 2022 Filippo A. E. Nuccio Mortarino Majno di Capriglio. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Filippo A. E. Nuccio, Junyan Xu +-/ +import topology.compact_open +import topology.homotopy.path + +/-! +# H-spaces + +This file defines H-spaces mainly following the approach proposed by Serre in his paper +*Homologie singulière des espaces fibrés*. The main results are the H-space `instance` on every +topological group, and the H-space structure on the loop space (based at `x : X` of any topological +space `X`, for which we introduce the notation `Ω_[x]`. + +## References + +* [J.-P. Serre, *Homologie singulière des espaces fibrés. Applications*, + Ann. of Math (2) 1951, 54, 425–505][serre1951] +-/ + +universes u v + +noncomputable theory + +open_locale unit_interval + +open path continuous_map set.Icc topological_space + +/-- +A topological space `X` is an H-space if it behaves like a (potentially non-associative) +topological group, but where the axioms for a group only hold up to homotopy. +-/ +class H_space (X : Type u) [topological_space X] := +(Hmul : C(X × X, X)) +(e : X) +(Hmul_e_e : Hmul (e, e) = e) +(e_Hmul : (Hmul.comp $ (const X e).prod_mk $ continuous_map.id X).homotopy_rel + (continuous_map.id X) {e}) +(Hmul_e : (Hmul.comp $ (continuous_map.id X).prod_mk $ const X e).homotopy_rel + (continuous_map.id X) {e}) + +/- We use the notation `⋀`, typeset as \And, to denote the binary operation `Hmul` on a H-space -/ +localized "notation (name := H_space.Hmul) x `⋀` y := H_space.Hmul (x, y) " in H_spaces + +instance H_space.prod (X : Type u) (Y : Type v) [topological_space X] [topological_space Y] +[H_space X] [H_space Y] : H_space (X × Y) := +{ Hmul := ⟨λ p, ((p.1.1 ⋀ p.2.1), p.1.2 ⋀ p.2.2), by continuity⟩, + e := (H_space.e, H_space.e), + Hmul_e_e := by {simp only [continuous_map.coe_mk, prod.mk.inj_iff], + exact ⟨H_space.Hmul_e_e, H_space.Hmul_e_e⟩}, + e_Hmul := + begin + let G : I × (X × Y) → X × Y := + (λ p, (H_space.e_Hmul (p.1, p.2.1), H_space.e_Hmul (p.1, p.2.2))), + have hG : continuous G := (continuous.comp H_space.e_Hmul.1.1.2 (continuous_fst.prod_mk + (continuous_fst.comp continuous_snd))).prod_mk (continuous.comp H_space.e_Hmul.1.1.2 + (continuous_fst.prod_mk (continuous_snd.comp continuous_snd))), + use ⟨G, hG⟩, + { rintros ⟨x, y⟩, + exacts prod.mk.inj_iff.mpr ⟨(H_space.e_Hmul).1.2 x, (H_space.e_Hmul).1.2 y⟩ }, + { rintros ⟨x, y⟩, + exact prod.mk.inj_iff.mpr ⟨(H_space.e_Hmul).1.3 x, (H_space.e_Hmul).1.3 y⟩ }, + { rintros t ⟨x, y⟩ h, + replace h := prod.mk.inj_iff.mp (set.mem_singleton_iff.mp h), + exact ⟨prod.mk.inj_iff.mpr ⟨homotopy_rel.eq_fst (H_space.e_Hmul) t + (set.mem_singleton_iff.mpr h.1), + homotopy_rel.eq_fst (H_space.e_Hmul) t (set.mem_singleton_iff.mpr h.2)⟩, + prod.mk.inj_iff.mpr ⟨((H_space.e_Hmul).2 t x h.1).2, ((H_space.e_Hmul).2 t y h.2).2⟩⟩ }, + end, + Hmul_e := + begin + let G : I × (X × Y) → X × Y := + (λ p, (H_space.Hmul_e (p.1, p.2.1), H_space.Hmul_e (p.1, p.2.2))), + have hG : continuous G := (continuous.comp H_space.Hmul_e.1.1.2 (continuous_fst.prod_mk + (continuous_fst.comp continuous_snd))).prod_mk (continuous.comp H_space.Hmul_e.1.1.2 + (continuous_fst.prod_mk (continuous_snd.comp continuous_snd))), + use ⟨G, hG⟩, + { rintros ⟨x, y⟩, + exacts prod.mk.inj_iff.mpr ⟨(H_space.Hmul_e).1.2 x, (H_space.Hmul_e).1.2 y⟩ }, + { rintros ⟨x, y⟩, + exact prod.mk.inj_iff.mpr ⟨(H_space.Hmul_e).1.3 x, (H_space.Hmul_e).1.3 y⟩ }, + { rintros t ⟨x, y⟩ h, + replace h := prod.mk.inj_iff.mp (set.mem_singleton_iff.mp h), + exact ⟨prod.mk.inj_iff.mpr ⟨homotopy_rel.eq_fst (H_space.Hmul_e) t + (set.mem_singleton_iff.mpr h.1), homotopy_rel.eq_fst (H_space.Hmul_e) t + (set.mem_singleton_iff.mpr h.2)⟩, prod.mk.inj_iff.mpr ⟨((H_space.Hmul_e).2 t x h.1).2, + ((H_space.Hmul_e).2 t y h.2).2⟩⟩ }, + end, } + +namespace topological_group + +@[priority 600, to_additive] instance H_space (G : Type u) [topological_space G] [group G] + [topological_group G] : H_space G := +{ Hmul := ⟨function.uncurry has_mul.mul, continuous_mul⟩, + e := 1, + Hmul_e_e := one_mul 1, + e_Hmul := (homotopy_rel.refl _ _).cast rfl (by {ext1, apply one_mul}), + Hmul_e := (homotopy_rel.refl _ _).cast rfl (by {ext1, apply mul_one}) } + +lemma one_eq_H_space_e {G : Type u} [topological_space G] [group G] [topological_group G] : + (1 : G) = H_space.e := rfl + +/- +In the following example we see that the `H-space` structure on the product of two topological +groups is definitionally equally to the product `H-space`-structure of the two groups. +-/ +example {G G' : Type u} [topological_space G] [group G] [topological_group G] + [topological_space G'] [group G'] [topological_group G'] : + topological_group.H_space (G × G') = H_space.prod G G' := rfl + + +end topological_group + +namespace unit_interval + +/-- `Q_right` is analogous to the function `Q` defined on p. 475 of [serre1951] that helps proving +continuity of `delay_refl_right`.-/ +def Q_right (p : I × I) : I := set.proj_Icc 0 1 zero_le_one (2 * p.1 / (1 + p.2)) + +lemma continuous_Q_right : continuous Q_right := +continuous_proj_Icc.comp $ continuous.div (by continuity) (by continuity) + (λ x, (add_pos zero_lt_one).ne') + +lemma Q_right_zero_left (θ : I) : Q_right (0, θ) = 0 := +set.proj_Icc_of_le_left _ $ by simp only [coe_zero, mul_zero, zero_div] + +lemma Q_right_one_left (θ : I) : Q_right (1, θ) = 1 := +set.proj_Icc_of_right_le _ $ (le_div_iff $ add_pos zero_lt_one).2 $ + by { dsimp only, rw [coe_one, one_mul, mul_one], apply add_le_add_left (le_one _) } + +lemma Q_right_zero_right (t : I) : (Q_right (t, 0) : ℝ) = if (t : ℝ) ≤ 1 / 2 then 2 * t else 1 := +begin + simp only [Q_right, coe_zero, add_zero, div_one], + split_ifs, + { rw set.proj_Icc_of_mem _ ((mul_pos_mem_iff zero_lt_two).2 _), exacts [rfl, ⟨t.2.1, h⟩] }, + { rw (set.proj_Icc_eq_right _).2, { refl }, { linarith }, { exact zero_lt_one } }, +end + +lemma Q_right_one_right (t : I) : Q_right (t, 1) = t := +eq.trans (by {rw Q_right, congr, apply mul_div_cancel_left, exact two_ne_zero}) $ + set.proj_Icc_coe zero_le_one _ + +end unit_interval + +namespace path + +open unit_interval + +variables {X : Type u} [topological_space X] {x y : X} + +/-- This is the function analogous to the one on p. 475 of [serre1951], defining a homotopy from +the product path `γ ∧ e` to `γ`.-/ +def delay_refl_right (θ : I) (γ : path x y) : path x y := +{ to_fun := λ t, γ (Q_right (t, θ)), + continuous_to_fun := γ.continuous.comp (continuous_Q_right.comp $ continuous.prod.mk_left θ), + source' := by { dsimp only, rw [Q_right_zero_left, γ.source] }, + target' := by { dsimp only, rw [Q_right_one_left, γ.target] } } + +lemma continuous_delay_refl_right : continuous (λ p : I × path x y, delay_refl_right p.1 p.2) := + continuous_uncurry_iff.mp $ (continuous_snd.comp continuous_fst).path_eval $ + continuous_Q_right.comp $ continuous_snd.prod_mk $ continuous_fst.comp continuous_fst + +lemma delay_refl_right_zero (γ : path x y) : delay_refl_right 0 γ = γ.trans (path.refl y) := +begin + ext t, + simp only [delay_refl_right, + trans_apply, refl_extend, path.coe_mk, function.comp_app, refl_apply], + split_ifs, swap, conv_rhs { rw ← γ.target }, + all_goals { apply congr_arg γ, ext1, rw Q_right_zero_right }, + exacts [if_neg h, if_pos h], +end + +lemma delay_refl_right_one (γ : path x y) : delay_refl_right 1 γ = γ := +by { ext t, exact congr_arg γ (Q_right_one_right t) } + +/-- This is the function on p. 475 of [serre1951], defining a homotopy from a path `γ` to the +product path `e ∧ γ`.-/ +def delay_refl_left (θ : I) (γ : path x y) : path x y := (delay_refl_right θ γ.symm).symm + +lemma continuous_delay_refl_left : continuous (λ p : I × path x y, delay_refl_left p.1 p.2) := +path.continuous_symm.comp $ continuous_delay_refl_right.comp $ continuous_fst.prod_mk $ + path.continuous_symm.comp continuous_snd + +lemma delay_refl_left_zero (γ : path x y) : delay_refl_left 0 γ = (path.refl x).trans γ := +by simp only [delay_refl_left, delay_refl_right_zero, trans_symm, refl_symm, path.symm_symm] + +lemma delay_refl_left_one (γ : path x y) : delay_refl_left 1 γ = γ := +by simp only [delay_refl_left, delay_refl_right_one, path.symm_symm] + +/-- +The loop space at x carries a structure of a `H-space`. Note that the field `e_Hmul` +(resp. `Hmul_e`) neither implies nor is implied by `path.homotopy.refl_trans` +(resp. `path.homotopy.trans_refl`). +-/ + +instance (x : X) : H_space (path x x) := +{ Hmul := ⟨λ ρ, ρ.1.trans ρ.2, continuous_trans⟩, + e := refl x, + Hmul_e_e := refl_trans_refl, + e_Hmul := + { to_homotopy := ⟨⟨λ p : I × (path x x), delay_refl_left p.1 p.2, + continuous_delay_refl_left⟩, delay_refl_left_zero, delay_refl_left_one⟩, + prop' := by { rintro t _ (rfl : _ = _), exact ⟨refl_trans_refl.symm, rfl⟩ } }, + Hmul_e := + { to_homotopy := ⟨⟨λ p : I × (path x x), delay_refl_right p.1 p.2, + continuous_delay_refl_right⟩, delay_refl_right_zero, delay_refl_right_one⟩, + prop' := by { rintro t _ (rfl : _ = _), exact ⟨refl_trans_refl.symm, rfl⟩ } } } + +end path From d4f691b9e5f94cfc64639973f3544c95f8d5d494 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Tue, 23 May 2023 21:45:53 +0000 Subject: [PATCH 1065/1291] feat(order/filter/basic): generalize some lemmas from `nhds_within` (#19070) --- src/order/filter/basic.lean | 13 +++++++++++++ src/topology/continuous_on.lean | 24 +++--------------------- 2 files changed, 16 insertions(+), 21 deletions(-) diff --git a/src/order/filter/basic.lean b/src/order/filter/basic.lean index 6748130b7d2ef..25f4b46a07e7b 100644 --- a/src/order/filter/basic.lean +++ b/src/order/filter/basic.lean @@ -1494,6 +1494,19 @@ h.mono $ λ x, mt (s \ s' : set α) ≤ᶠ[l] (t \ t' : set α) := h.inter h'.compl +lemma set_eventually_le_iff_mem_inf_principal {s t : set α} {l : filter α} : + s ≤ᶠ[l] t ↔ t ∈ l ⊓ 𝓟 s := +mem_inf_principal.symm + +lemma set_eventually_le_iff_inf_principal_le {s t : set α} {l : filter α} : + s ≤ᶠ[l] t ↔ l ⊓ 𝓟 s ≤ l ⊓ 𝓟 t := +set_eventually_le_iff_mem_inf_principal.trans $ + by simp only [le_inf_iff, inf_le_left, true_and, le_principal_iff] + +lemma set_eventually_eq_iff_inf_principal {s t : set α} {l : filter α} : + s =ᶠ[l] t ↔ l ⊓ 𝓟 s = l ⊓ 𝓟 t := +by simp only [eventually_le_antisymm_iff, le_antisymm_iff, set_eventually_le_iff_inf_principal_le] + lemma eventually_le.mul_le_mul [mul_zero_class β] [partial_order β] [pos_mul_mono β] [mul_pos_mono β] {l : filter α} {f₁ f₂ g₁ g₂ : α → β} diff --git a/src/topology/continuous_on.lean b/src/topology/continuous_on.lean index bec903efb4b3f..8cb6fabef4bba 100644 --- a/src/topology/continuous_on.lean +++ b/src/topology/continuous_on.lean @@ -107,35 +107,17 @@ end lemma mem_nhds_within_iff_eventually {s t : set α} {x : α} : t ∈ 𝓝[s] x ↔ ∀ᶠ y in 𝓝 x, y ∈ s → y ∈ t := -begin - rw [mem_nhds_within_iff_exists_mem_nhds_inter], - split, - { rintro ⟨u, hu, hut⟩, exact eventually_of_mem hu (λ x hxu hxs, hut ⟨hxu, hxs⟩) }, - { refine λ h, ⟨_, h, λ y hy, hy.1 hy.2⟩ } -end +set_eventually_le_iff_mem_inf_principal.symm lemma mem_nhds_within_iff_eventually_eq {s t : set α} {x : α} : t ∈ 𝓝[s] x ↔ s =ᶠ[𝓝 x] (s ∩ t : set α) := by simp_rw [mem_nhds_within_iff_eventually, eventually_eq_set, mem_inter_iff, iff_self_and] lemma nhds_within_eq_iff_eventually_eq {s t : set α} {x : α} : 𝓝[s] x = 𝓝[t] x ↔ s =ᶠ[𝓝 x] t := -begin - simp_rw [filter.ext_iff, mem_nhds_within_iff_eventually, eventually_eq_set], - split, - { intro h, - filter_upwards [(h t).mpr (eventually_of_forall $ λ x, id), - (h s).mp (eventually_of_forall $ λ x, id)], - exact λ x, iff.intro, }, - { refine λ h u, eventually_congr (h.mono $ λ x h, _), rw [h] } -end +set_eventually_eq_iff_inf_principal.symm lemma nhds_within_le_iff {s t : set α} {x : α} : 𝓝[s] x ≤ 𝓝[t] x ↔ t ∈ 𝓝[s] x := -begin - simp_rw [filter.le_def, mem_nhds_within_iff_eventually], - split, - { exact λ h, (h t $ eventually_of_forall (λ x, id)).mono (λ x, id) }, - { exact λ h u hu, (h.and hu).mono (λ x hx h, hx.2 $ hx.1 h) } -end +set_eventually_le_iff_inf_principal_le.symm.trans set_eventually_le_iff_mem_inf_principal lemma preimage_nhds_within_coinduced' {π : α → β} {s : set β} {t : set α} {a : α} (h : a ∈ t) (ht : is_open t) From 3a69562db5a458db8322b190ec8d9a8bbd8a5b14 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Wed, 24 May 2023 02:31:04 +0000 Subject: [PATCH 1066/1291] feat(analysis/calculus): drop unneeded assumptions (#19045) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit * Prove more precise congruence lemmas for derivatives. In particular, adding or removing a single point from `s` doesn't change anything about derivatives within `s`. * Drop `has_fderiv_within_at.antimono` and `has_deriv_within_at.antimono`, use stronger `*.mono_of_mem` lemmas instead. * Prove some equalities about `(f)deriv_within` by using `simp` instead of `slit_ifs`: if `has_fderiv_within_at f f' s x ↔ has_fderiv_within_at g g' t y`, then `fderiv_within f s x = fderiv_within g t y`. --- src/analysis/calculus/cont_diff.lean | 35 ++-- src/analysis/calculus/cont_diff_def.lean | 181 +++++++++------- src/analysis/calculus/deriv.lean | 44 ++-- src/analysis/calculus/fderiv/basic.lean | 196 +++++++++--------- src/analysis/calculus/iterated_deriv.lean | 17 +- src/analysis/calculus/taylor.lean | 75 +++---- src/geometry/manifold/cont_mdiff_mfderiv.lean | 12 +- src/geometry/manifold/mfderiv.lean | 4 +- .../manifold/vector_bundle/tangent.lean | 2 - 9 files changed, 295 insertions(+), 271 deletions(-) diff --git a/src/analysis/calculus/cont_diff.lean b/src/analysis/calculus/cont_diff.lean index 2a8cced40e013..dd1f41175b230 100644 --- a/src/analysis/calculus/cont_diff.lean +++ b/src/analysis/calculus/cont_diff.lean @@ -327,7 +327,7 @@ begin have Z : fderiv_within 𝕜 (iterated_fderiv_within 𝕜 i (g ∘ f) s) s x = fderiv_within 𝕜 (λ y, g.comp_continuous_multilinear_mapL (λ (j : fin i), E) (iterated_fderiv_within 𝕜 i f s y)) s x, - from fderiv_within_congr' (hs x hx) (λ y hy, IH hy) hx, + from fderiv_within_congr' @IH hx, simp_rw Z, rw (g.comp_continuous_multilinear_mapL (λ (j : fin i), E)).comp_fderiv_within (hs x hx), simp only [continuous_linear_map.coe_comp', continuous_linear_equiv.coe_coe, comp_app, @@ -495,7 +495,7 @@ begin have : fderiv_within 𝕜 (iterated_fderiv_within 𝕜 i (f ∘ ⇑g) (⇑g ⁻¹' s)) (⇑g ⁻¹' s) x = fderiv_within 𝕜 (λ y, continuous_multilinear_map.comp_continuous_linear_map_equivL _ (λ (_x : fin i), g) (iterated_fderiv_within 𝕜 i f s (g y))) (g ⁻¹' s) x, - from fderiv_within_congr' (g.unique_diff_on_preimage_iff.2 hs x hx) (λ y hy, IH hy) hx, + from fderiv_within_congr' @IH hx, rw [this], rw continuous_linear_equiv.comp_fderiv_within _ (g.unique_diff_on_preimage_iff.2 hs x hx), simp only [continuous_linear_map.coe_comp', continuous_linear_equiv.coe_coe, comp_app, @@ -1332,14 +1332,10 @@ begin = fderiv_within 𝕜 (iterated_fderiv_within 𝕜 i (f + g) s) s x (h 0) (fin.tail h) : rfl ... = fderiv_within 𝕜 (iterated_fderiv_within 𝕜 i f s + iterated_fderiv_within 𝕜 i g s) s x (h 0) (fin.tail h) : - begin - congr' 2, - exact fderiv_within_congr (hu x hx) (λ _, hi hcdf hcdg) (hi hcdf hcdg hx), - end + by { rw [fderiv_within_congr' (λ _, hi hcdf hcdg) hx], refl } ... = (fderiv_within 𝕜 (iterated_fderiv_within 𝕜 i f s) s + - fderiv_within 𝕜 (iterated_fderiv_within 𝕜 i g s) s) - x (h 0) (fin.tail h) : - by rw [pi.add_def, fderiv_within_add (hu x hx) (hdf x hx) (hdg x hx)]; refl + fderiv_within 𝕜 (iterated_fderiv_within 𝕜 i g s) s) x (h 0) (fin.tail h) : + by { rw [pi.add_def, fderiv_within_add (hu x hx) (hdf x hx) (hdg x hx)], refl } ... = (iterated_fderiv_within 𝕜 (i+1) f s + iterated_fderiv_within 𝕜 (i+1) g s) x h : rfl } end @@ -1409,14 +1405,10 @@ begin with_top.coe_lt_coe.mpr (nat.lt_succ_self _), calc iterated_fderiv_within 𝕜 (i+1) (-f) s x h = fderiv_within 𝕜 (iterated_fderiv_within 𝕜 i (-f) s) s x (h 0) (fin.tail h) : rfl - ... = fderiv_within 𝕜 (-iterated_fderiv_within 𝕜 i f s) s x - (h 0) (fin.tail h) : - begin - congr' 2, - exact fderiv_within_congr (hu x hx) (λ _, hi) (hi hx), - end + ... = fderiv_within 𝕜 (-iterated_fderiv_within 𝕜 i f s) s x (h 0) (fin.tail h) : + by { rw [fderiv_within_congr' @hi hx], refl } ... = -(fderiv_within 𝕜 (iterated_fderiv_within 𝕜 i f s) s) x (h 0) (fin.tail h) : - by rw [pi.neg_def, fderiv_within_neg (hu x hx)]; refl + by { rw [pi.neg_def, fderiv_within_neg (hu x hx)], refl } ... = - (iterated_fderiv_within 𝕜 (i+1) f s) x h : rfl } end @@ -1675,12 +1667,9 @@ begin calc iterated_fderiv_within 𝕜 (i+1) (a • f) s x h = fderiv_within 𝕜 (iterated_fderiv_within 𝕜 i (a • f) s) s x (h 0) (fin.tail h) : rfl ... = fderiv_within 𝕜 (a • iterated_fderiv_within 𝕜 i f s) s x (h 0) (fin.tail h) : - begin - congr' 2, - exact fderiv_within_congr (hu x hx) (λ _, hi hcdf) (hi hcdf hx), - end + by { rw [fderiv_within_congr' (λ _, hi hcdf) hx], refl } ... = (a • fderiv_within 𝕜 (iterated_fderiv_within 𝕜 i f s)) s x (h 0) (fin.tail h) : - by rw [pi.smul_def, fderiv_within_const_smul (hu x hx) (hdf x hx)]; refl + by { rw [pi.smul_def, fderiv_within_const_smul (hu x hx) (hdf x hx)], refl } ... = a • iterated_fderiv_within 𝕜 (i+1) f s x h : rfl } end @@ -2432,7 +2421,7 @@ begin (λ (y : Du), fderiv_within 𝕜 (λ (y : Du), B (f y) (g y)) s y) s x = iterated_fderiv_within 𝕜 n (λ y, B.precompR Du (f y) (fderiv_within 𝕜 g s y) + B.precompL Du (fderiv_within 𝕜 f s y) (g y)) s x, - { apply iterated_fderiv_within_congr hs (λ y hy, _) hx, + { apply iterated_fderiv_within_congr (λ y hy, _) hx, have L : (1 : ℕ∞) ≤ n.succ, by simpa only [enat.coe_one, nat.one_le_cast] using nat.succ_pos n, exact B.fderiv_within_of_bilinear (hf.differentiable_on L y hy) @@ -2707,7 +2696,7 @@ begin begin have L : (1 : ℕ∞) ≤ n.succ, by simpa only [enat.coe_one, nat.one_le_cast] using n.succ_pos, congr' 1, - apply iterated_fderiv_within_congr hs (λ y hy, _) hx, + refine iterated_fderiv_within_congr (λ y hy, _) hx _, apply fderiv_within.comp _ _ _ hst (hs y hy), { exact hg.differentiable_on L _ (hst hy) }, { exact hf.differentiable_on L _ hy } diff --git a/src/analysis/calculus/cont_diff_def.lean b/src/analysis/calculus/cont_diff_def.lean index 8e64de8d50d92..80e603f1b05f7 100644 --- a/src/analysis/calculus/cont_diff_def.lean +++ b/src/analysis/calculus/cont_diff_def.lean @@ -156,7 +156,7 @@ derivative, differentiability, higher derivative, `C^n`, multilinear, Taylor ser -/ noncomputable theory -open_locale classical big_operators nnreal topology +open_locale classical big_operators nnreal topology filter local notation `∞` := (⊤ : ℕ∞) @@ -645,19 +645,22 @@ lemma cont_diff_on.cont_diff_within_at (h : cont_diff_on 𝕜 n f s) (hx : x ∈ cont_diff_within_at 𝕜 n f s x := h x hx -lemma cont_diff_within_at.cont_diff_on {m : ℕ} +lemma cont_diff_within_at.cont_diff_on' {m : ℕ} (hm : (m : ℕ∞) ≤ n) (h : cont_diff_within_at 𝕜 n f s x) : - ∃ u ∈ 𝓝[insert x s] x, u ⊆ insert x s ∧ cont_diff_on 𝕜 m f u := + ∃ u, is_open u ∧ x ∈ u ∧ cont_diff_on 𝕜 m f (insert x s ∩ u) := begin - rcases h m hm with ⟨u, u_nhd, p, hp⟩, - refine ⟨u ∩ insert x s, filter.inter_mem u_nhd self_mem_nhds_within, - inter_subset_right _ _, _⟩, - assume y hy m' hm', - refine ⟨u ∩ insert x s, _, p, (hp.mono (inter_subset_left _ _)).of_le hm'⟩, - convert self_mem_nhds_within, - exact insert_eq_of_mem hy + rcases h m hm with ⟨t, ht, p, hp⟩, + rcases mem_nhds_within.1 ht with ⟨u, huo, hxu, hut⟩, + rw [inter_comm] at hut, + exact ⟨u, huo, hxu, (hp.mono hut).cont_diff_on⟩ end +lemma cont_diff_within_at.cont_diff_on {m : ℕ} + (hm : (m : ℕ∞) ≤ n) (h : cont_diff_within_at 𝕜 n f s x) : + ∃ u ∈ 𝓝[insert x s] x, u ⊆ insert x s ∧ cont_diff_on 𝕜 m f u := +let ⟨u, uo, xu, h⟩ := h.cont_diff_on' hm +in ⟨_, inter_mem_nhds_within _ (uo.mem_nhds xu), inter_subset_left _ _, h⟩ + protected lemma cont_diff_within_at.eventually {n : ℕ} (h : cont_diff_within_at 𝕜 n f s x) : ∀ᶠ y in 𝓝[insert x s] x, cont_diff_within_at 𝕜 n f s y := @@ -824,7 +827,7 @@ begin : E → (E [×(n + 1)]→L[𝕜] F)) (m 0) (tail m) : rfl ... = (fderiv_within 𝕜 (I ∘ (iterated_fderiv_within 𝕜 n (fderiv_within 𝕜 f s) s)) s x : E → (E [×(n + 1)]→L[𝕜] F)) (m 0) (tail m) : - by rw fderiv_within_congr (hs x hx) A (A x hx) + by rw fderiv_within_congr A (A x hx) ... = (I ∘ fderiv_within 𝕜 ((iterated_fderiv_within 𝕜 n (fderiv_within 𝕜 f s) s)) s x : E → (E [×(n + 1)]→L[𝕜] F)) (m 0) (tail m) : by { rw linear_isometry_equiv.comp_fderiv_within _ (hs x hx), refl } @@ -849,72 +852,87 @@ lemma norm_iterated_fderiv_within_fderiv_within {n : ℕ} (hs : unique_diff_on by rw [iterated_fderiv_within_succ_eq_comp_right hs hx, linear_isometry_equiv.norm_map] @[simp] lemma iterated_fderiv_within_one_apply - (hs : unique_diff_on 𝕜 s) (hx : x ∈ s) (m : (fin 1) → E) : + (h : unique_diff_within_at 𝕜 s x) (m : fin 1 → E) : (iterated_fderiv_within 𝕜 1 f s x : ((fin 1) → E) → F) m = (fderiv_within 𝕜 f s x : E → F) (m 0) := -by { rw [iterated_fderiv_within_succ_apply_right hs hx, iterated_fderiv_within_zero_apply], refl } +begin + simp only [iterated_fderiv_within_succ_apply_left, iterated_fderiv_within_zero_eq_comp, + (continuous_multilinear_curry_fin0 𝕜 E F).symm.comp_fderiv_within h], + refl +end -/-- If two functions coincide on a set `s` of unique differentiability, then their iterated -differentials within this set coincide. -/ -lemma iterated_fderiv_within_congr {n : ℕ} - (hs : unique_diff_on 𝕜 s) (hL : ∀y∈s, f₁ y = f y) (hx : x ∈ s) : - iterated_fderiv_within 𝕜 n f₁ s x = iterated_fderiv_within 𝕜 n f s x := +lemma filter.eventually_eq.iterated_fderiv_within' (h : f₁ =ᶠ[𝓝[s] x] f) (ht : t ⊆ s) (n : ℕ) : + iterated_fderiv_within 𝕜 n f₁ t =ᶠ[𝓝[s] x] iterated_fderiv_within 𝕜 n f t := begin - induction n with n IH generalizing x, - { ext m, simp [hL x hx] }, - { have : fderiv_within 𝕜 (λ y, iterated_fderiv_within 𝕜 n f₁ s y) s x - = fderiv_within 𝕜 (λ y, iterated_fderiv_within 𝕜 n f s y) s x := - fderiv_within_congr (hs x hx) (λ y hy, IH hy) (IH hx), - ext m, - rw [iterated_fderiv_within_succ_apply_left, iterated_fderiv_within_succ_apply_left, this] } + induction n with n ihn, + { exact h.mono (λ y hy, fun_like.ext _ _ $ λ _, hy) }, + { have : fderiv_within 𝕜 _ t =ᶠ[𝓝[s] x] fderiv_within 𝕜 _ t := ihn.fderiv_within' ht, + apply this.mono, + intros y hy, + simp only [iterated_fderiv_within_succ_eq_comp_left, hy, (∘)] } end -/-- The iterated differential within a set `s` at a point `x` is not modified if one intersects -`s` with an open set containing `x`. -/ -lemma iterated_fderiv_within_inter_open {n : ℕ} (hu : is_open u) - (hs : unique_diff_on 𝕜 (s ∩ u)) (hx : x ∈ s ∩ u) : - iterated_fderiv_within 𝕜 n f (s ∩ u) x = iterated_fderiv_within 𝕜 n f s x := +protected lemma filter.eventually_eq.iterated_fderiv_within (h : f₁ =ᶠ[𝓝[s] x] f) (n : ℕ) : + iterated_fderiv_within 𝕜 n f₁ s =ᶠ[𝓝[s] x] iterated_fderiv_within 𝕜 n f s := +h.iterated_fderiv_within' subset.rfl n + +/-- If two functions coincide in a neighborhood of `x` within a set `s` and at `x`, then their +iterated differentials within this set at `x` coincide. -/ +lemma filter.eventually_eq.iterated_fderiv_within_eq (h : f₁ =ᶠ[𝓝[s] x] f) (hx : f₁ x = f x) + (n : ℕ) : iterated_fderiv_within 𝕜 n f₁ s x = iterated_fderiv_within 𝕜 n f s x := +have f₁ =ᶠ[𝓝[insert x s] x] f, by simpa [eventually_eq, hx], +(this.iterated_fderiv_within' (subset_insert _ _) n).self_of_nhds_within (mem_insert _ _) + +/-- If two functions coincide on a set `s`, then their iterated differentials within this set +coincide. See also `filter.eventually_eq.iterated_fderiv_within_eq` and +`filter.eventually_eq.iterated_fderiv_within`. -/ +lemma iterated_fderiv_within_congr (hs : eq_on f₁ f s) (hx : x ∈ s) (n : ℕ) : + iterated_fderiv_within 𝕜 n f₁ s x = iterated_fderiv_within 𝕜 n f s x := +(hs.eventually_eq.filter_mono inf_le_right).iterated_fderiv_within_eq (hs hx) _ + +/-- If two functions coincide on a set `s`, then their iterated differentials within this set +coincide. See also `filter.eventually_eq.iterated_fderiv_within_eq` and +`filter.eventually_eq.iterated_fderiv_within`. -/ +protected lemma set.eq_on.iterated_fderiv_within (hs : eq_on f₁ f s) (n : ℕ) : + eq_on (iterated_fderiv_within 𝕜 n f₁ s) (iterated_fderiv_within 𝕜 n f s) s := +λ x hx, iterated_fderiv_within_congr hs hx n + +lemma iterated_fderiv_within_eventually_congr_set' (y : E) (h : s =ᶠ[𝓝[{y}ᶜ] x] t) (n : ℕ) : + iterated_fderiv_within 𝕜 n f s =ᶠ[𝓝 x] iterated_fderiv_within 𝕜 n f t := begin - induction n with n IH generalizing x, - { ext m, simp }, - { have A : fderiv_within 𝕜 (λ y, iterated_fderiv_within 𝕜 n f (s ∩ u) y) (s ∩ u) x - = fderiv_within 𝕜 (λ y, iterated_fderiv_within 𝕜 n f s y) (s ∩ u) x := - fderiv_within_congr (hs x hx) (λ y hy, IH hy) (IH hx), - have B : fderiv_within 𝕜 (λ y, iterated_fderiv_within 𝕜 n f s y) (s ∩ u) x - = fderiv_within 𝕜 (λ y, iterated_fderiv_within 𝕜 n f s y) s x := - fderiv_within_inter (is_open.mem_nhds hu hx.2) - ((unique_diff_within_at_inter (is_open.mem_nhds hu hx.2)).1 (hs x hx)), - ext m, - rw [iterated_fderiv_within_succ_apply_left, iterated_fderiv_within_succ_apply_left, A, B] } + induction n with n ihn generalizing x, + { refl }, + { refine (eventually_nhds_nhds_within.2 h).mono (λ y hy, _), + simp only [iterated_fderiv_within_succ_eq_comp_left, (∘)], + rw [(ihn hy).fderiv_within_eq_nhds, fderiv_within_congr_set' _ hy] } end +lemma iterated_fderiv_within_eventually_congr_set (h : s =ᶠ[𝓝 x] t) (n : ℕ) : + iterated_fderiv_within 𝕜 n f s =ᶠ[𝓝 x] iterated_fderiv_within 𝕜 n f t := +iterated_fderiv_within_eventually_congr_set' x (h.filter_mono inf_le_left) n + +lemma iterated_fderiv_within_congr_set (h : s =ᶠ[𝓝 x] t) (n : ℕ) : + iterated_fderiv_within 𝕜 n f s x = iterated_fderiv_within 𝕜 n f t x := +(iterated_fderiv_within_eventually_congr_set h n).self_of_nhds + /-- The iterated differential within a set `s` at a point `x` is not modified if one intersects `s` with a neighborhood of `x` within `s`. -/ -lemma iterated_fderiv_within_inter' {n : ℕ} - (hu : u ∈ 𝓝[s] x) (hs : unique_diff_on 𝕜 s) (xs : x ∈ s) : +lemma iterated_fderiv_within_inter' {n : ℕ} (hu : u ∈ 𝓝[s] x) : iterated_fderiv_within 𝕜 n f (s ∩ u) x = iterated_fderiv_within 𝕜 n f s x := -begin - obtain ⟨v, v_open, xv, vu⟩ : ∃ v, is_open v ∧ x ∈ v ∧ v ∩ s ⊆ u := mem_nhds_within.1 hu, - have A : (s ∩ u) ∩ v = s ∩ v, - { apply subset.antisymm (inter_subset_inter (inter_subset_left _ _) (subset.refl _)), - exact λ y ⟨ys, yv⟩, ⟨⟨ys, vu ⟨yv, ys⟩⟩, yv⟩ }, - have : iterated_fderiv_within 𝕜 n f (s ∩ v) x = iterated_fderiv_within 𝕜 n f s x := - iterated_fderiv_within_inter_open v_open (hs.inter v_open) ⟨xs, xv⟩, - rw ← this, - have : iterated_fderiv_within 𝕜 n f ((s ∩ u) ∩ v) x = iterated_fderiv_within 𝕜 n f (s ∩ u) x, - { refine iterated_fderiv_within_inter_open v_open _ ⟨⟨xs, vu ⟨xv, xs⟩⟩, xv⟩, - rw A, - exact hs.inter v_open }, - rw A at this, - rw ← this -end +iterated_fderiv_within_congr_set (nhds_within_eq_iff_eventually_eq.1 $ nhds_within_inter_of_mem' hu) + _ /-- The iterated differential within a set `s` at a point `x` is not modified if one intersects `s` with a neighborhood of `x`. -/ -lemma iterated_fderiv_within_inter {n : ℕ} - (hu : u ∈ 𝓝 x) (hs : unique_diff_on 𝕜 s) (xs : x ∈ s) : +lemma iterated_fderiv_within_inter {n : ℕ} (hu : u ∈ 𝓝 x) : iterated_fderiv_within 𝕜 n f (s ∩ u) x = iterated_fderiv_within 𝕜 n f s x := -iterated_fderiv_within_inter' (mem_nhds_within_of_mem_nhds hu) hs xs +iterated_fderiv_within_inter' (mem_nhds_within_of_mem_nhds hu) + +/-- The iterated differential within a set `s` at a point `x` is not modified if one intersects +`s` with an open set containing `x`. -/ +lemma iterated_fderiv_within_inter_open {n : ℕ} (hu : is_open u) (hx : x ∈ u) : + iterated_fderiv_within 𝕜 n f (s ∩ u) x = iterated_fderiv_within 𝕜 n f s x := +iterated_fderiv_within_inter (hu.mem_nhds hx) @[simp] lemma cont_diff_on_zero : cont_diff_on 𝕜 0 f s ↔ continuous_on f s := @@ -979,14 +997,14 @@ begin rw inter_comm at ho, have : p x m.succ = ftaylor_series_within 𝕜 f s x m.succ, { change p x m.succ = iterated_fderiv_within 𝕜 m.succ f s x, - rw ← iterated_fderiv_within_inter (is_open.mem_nhds o_open xo) hs hx, + rw [← iterated_fderiv_within_inter_open o_open xo], exact (Hp.mono ho).eq_ftaylor_series_of_unique_diff_on le_rfl (hs.inter o_open) ⟨hx, xo⟩ }, rw [← this, ← has_fderiv_within_at_inter (is_open.mem_nhds o_open xo)], have A : ∀ y ∈ s ∩ o, p y m = ftaylor_series_within 𝕜 f s y m, { rintros y ⟨hy, yo⟩, change p y m = iterated_fderiv_within 𝕜 m f s y, - rw ← iterated_fderiv_within_inter (is_open.mem_nhds o_open yo) hs hy, + rw [← iterated_fderiv_within_inter_open o_open yo], exact (Hp.mono ho).eq_ftaylor_series_of_unique_diff_on (with_top.coe_le_coe.2 (nat.le_succ m)) (hs.inter o_open) ⟨hy, yo⟩ }, exact ((Hp.mono ho).fderiv_within m (with_top.coe_lt_coe.2 (lt_add_one m)) x ⟨hx, xo⟩).congr @@ -1002,7 +1020,7 @@ begin have A : ∀ y ∈ s ∩ o, p y m = ftaylor_series_within 𝕜 f s y m, { rintros y ⟨hy, yo⟩, change p y m = iterated_fderiv_within 𝕜 m f s y, - rw ← iterated_fderiv_within_inter (is_open.mem_nhds o_open yo) hs hy, + rw [← iterated_fderiv_within_inter_open o_open yo], exact (Hp.mono ho).eq_ftaylor_series_of_unique_diff_on le_rfl (hs.inter o_open) ⟨hy, yo⟩ }, exact ((Hp.mono ho).cont m le_rfl).congr (λ y hy, (A y hy).symm) } @@ -1047,6 +1065,27 @@ lemma cont_diff_on.differentiable_on_iterated_fderiv_within {m : ℕ} differentiable_on 𝕜 (iterated_fderiv_within 𝕜 m f s) s := λ x hx, ((h.ftaylor_series_within hs).fderiv_within m hmn x hx).differentiable_within_at +lemma cont_diff_within_at.differentiable_within_at_iterated_fderiv_within {m : ℕ} + (h : cont_diff_within_at 𝕜 n f s x) (hmn : (m : ℕ∞) < n) + (hs : unique_diff_on 𝕜 (insert x s)) : + differentiable_within_at 𝕜 (iterated_fderiv_within 𝕜 m f s) s x := +begin + rcases h.cont_diff_on' (enat.add_one_le_of_lt hmn) with ⟨u, uo, xu, hu⟩, + set t := insert x s ∩ u, + have A : t =ᶠ[𝓝[≠] x] s, + { simp only [set_eventually_eq_iff_inf_principal, ← nhds_within_inter'], + rw [← inter_assoc, nhds_within_inter_of_mem', ← diff_eq_compl_inter, insert_diff_of_mem, + diff_eq_compl_inter], + exacts [rfl, mem_nhds_within_of_mem_nhds (uo.mem_nhds xu)] }, + have B : iterated_fderiv_within 𝕜 m f s =ᶠ[𝓝 x] iterated_fderiv_within 𝕜 m f t, + from iterated_fderiv_within_eventually_congr_set' _ A.symm _, + have C : differentiable_within_at 𝕜 (iterated_fderiv_within 𝕜 m f t) t x, + from hu.differentiable_on_iterated_fderiv_within (nat.cast_lt.2 m.lt_succ_self) (hs.inter uo) x + ⟨mem_insert _ _, xu⟩, + rw [differentiable_within_at_congr_set' _ A] at C, + exact C.congr_of_eventually_eq (B.filter_mono inf_le_left) B.self_of_nhds +end + lemma cont_diff_on_iff_continuous_on_differentiable_on (hs : unique_diff_on 𝕜 s) : cont_diff_on 𝕜 n f s ↔ @@ -1054,15 +1093,9 @@ lemma cont_diff_on_iff_continuous_on_differentiable_on continuous_on (λ x, iterated_fderiv_within 𝕜 m f s x) s) ∧ (∀ (m : ℕ), (m : ℕ∞) < n → differentiable_on 𝕜 (λ x, iterated_fderiv_within 𝕜 m f s x) s) := -begin - split, - { assume h, - split, - { assume m hm, exact h.continuous_on_iterated_fderiv_within hm hs }, - { assume m hm, exact h.differentiable_on_iterated_fderiv_within hm hs } }, - { assume h, - exact cont_diff_on_of_continuous_on_differentiable_on h.1 h.2 } -end +⟨λ h, ⟨λ m hm, h.continuous_on_iterated_fderiv_within hm hs, + λ m hm, h.differentiable_on_iterated_fderiv_within hm hs⟩, + λ h, cont_diff_on_of_continuous_on_differentiable_on h.1 h.2⟩ 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) : @@ -1095,7 +1128,7 @@ begin 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 + rwa fderiv_within_inter (o_open.mem_nhds hy.2) at A end lemma cont_diff_on_succ_iff_has_fderiv_within {n : ℕ} (hs : unique_diff_on 𝕜 s) : diff --git a/src/analysis/calculus/deriv.lean b/src/analysis/calculus/deriv.lean index 60f5e374fb9d7..dbcd04015b9bb 100644 --- a/src/analysis/calculus/deriv.lean +++ b/src/analysis/calculus/deriv.lean @@ -294,16 +294,19 @@ lemma has_deriv_at_iff_tendsto_slope : has_deriv_at f f' x ↔ tendsto (slope f x) (𝓝[≠] x) (𝓝 f') := has_deriv_at_filter_iff_tendsto_slope -theorem has_deriv_within_at_congr_set {s t u : set 𝕜} - (hu : u ∈ 𝓝 x) (h : s ∩ u = t ∩ u) : +theorem has_deriv_within_at_congr_set' {s t : set 𝕜} (y : 𝕜) (h : s =ᶠ[𝓝[{y}ᶜ] x] t) : has_deriv_within_at f f' s x ↔ has_deriv_within_at f f' t x := -by simp_rw [has_deriv_within_at, nhds_within_eq_nhds_within' hu h] +has_fderiv_within_at_congr_set' y h + +theorem has_deriv_within_at_congr_set {s t : set 𝕜} (h : s =ᶠ[𝓝 x] t) : + has_deriv_within_at f f' s x ↔ has_deriv_within_at f f' t x := +has_fderiv_within_at_congr_set h alias has_deriv_within_at_congr_set ↔ has_deriv_within_at.congr_set _ @[simp] lemma has_deriv_within_at_diff_singleton : has_deriv_within_at f f' (s \ {x}) x ↔ has_deriv_within_at f f' s x := -by simp only [has_deriv_within_at_iff_tendsto_slope, sdiff_idem] +has_fderiv_within_at_diff_singleton _ @[simp] lemma has_deriv_within_at_Ioi_iff_Ici [partial_order 𝕜] : has_deriv_within_at f f' (Ioi x) x ↔ has_deriv_within_at f f' (Ici x) x := @@ -322,8 +325,7 @@ alias has_deriv_within_at_Iio_iff_Iic ↔ theorem has_deriv_within_at.Ioi_iff_Ioo [linear_order 𝕜] [order_closed_topology 𝕜] {x y : 𝕜} (h : x < y) : has_deriv_within_at f f' (Ioo x y) x ↔ has_deriv_within_at f f' (Ioi x) x := -has_deriv_within_at_congr_set (is_open_Iio.mem_nhds h) $ - by { rw [Ioi_inter_Iio, inter_eq_left_iff_subset], exact Ioo_subset_Iio_self } +has_fderiv_within_at_inter $ Iio_mem_nhds h alias has_deriv_within_at.Ioi_iff_Ioo ↔ has_deriv_within_at.Ioi_of_Ioo has_deriv_within_at.Ioo_of_Ioi @@ -340,6 +342,10 @@ theorem has_deriv_within_at.mono (h : has_deriv_within_at f f' t x) (hst : s ⊆ has_deriv_within_at f f' s x := has_fderiv_within_at.mono h hst +theorem has_deriv_within_at.mono_of_mem (h : has_deriv_within_at f f' t x) (hst : t ∈ 𝓝[s] x) : + has_deriv_within_at f f' s x := +has_fderiv_within_at.mono_of_mem h hst + theorem has_deriv_at.has_deriv_at_filter (h : has_deriv_at f f' x) (hL : L ≤ 𝓝 x) : has_deriv_at_filter f f' x L := has_fderiv_at.has_fderiv_at_filter h hL @@ -435,17 +441,29 @@ theorem has_deriv_within_at.deriv_eq_zero (hd : has_deriv_within_at f 0 s x) (em' (differentiable_at 𝕜 f x)).elim deriv_zero_of_not_differentiable_at $ λ h, H.eq_deriv _ h.has_deriv_at.has_deriv_within_at hd +lemma deriv_within_of_mem (st : t ∈ 𝓝[s] x) (ht : unique_diff_within_at 𝕜 s x) + (h : differentiable_within_at 𝕜 f t x) : + deriv_within f s x = deriv_within f t x := +((differentiable_within_at.has_deriv_within_at h).mono_of_mem st).deriv_within ht + lemma deriv_within_subset (st : s ⊆ t) (ht : unique_diff_within_at 𝕜 s x) (h : differentiable_within_at 𝕜 f t x) : deriv_within f s x = deriv_within f t x := ((differentiable_within_at.has_deriv_within_at h).mono st).deriv_within ht +lemma deriv_within_congr_set' (y : 𝕜) (h : s =ᶠ[𝓝[{y}ᶜ] x] t) : + deriv_within f s x = deriv_within f t x := +by simp only [deriv_within, fderiv_within_congr_set' y h] + +lemma deriv_within_congr_set (h : s =ᶠ[𝓝 x] t) : deriv_within f s x = deriv_within f t x := +by simp only [deriv_within, fderiv_within_congr_set h] + @[simp] lemma deriv_within_univ : deriv_within f univ = deriv f := by { ext, unfold deriv_within deriv, rw fderiv_within_univ } -lemma deriv_within_inter (ht : t ∈ 𝓝 x) (hs : unique_diff_within_at 𝕜 s x) : +lemma deriv_within_inter (ht : t ∈ 𝓝 x) : deriv_within f (s ∩ t) x = deriv_within f s x := -by { unfold deriv_within, rw fderiv_within_inter ht hs } +by { unfold deriv_within, rw fderiv_within_inter ht } lemma deriv_within_of_open (hs : is_open s) (hx : x ∈ s) : deriv_within f s x = deriv f x := @@ -516,15 +534,13 @@ lemma has_deriv_at.congr_of_eventually_eq (h : has_deriv_at f f' x) (h₁ : f₁ =ᶠ[𝓝 x] f) : has_deriv_at f₁ f' x := has_deriv_at_filter.congr_of_eventually_eq h h₁ (mem_of_mem_nhds h₁ : _) -lemma filter.eventually_eq.deriv_within_eq (hs : unique_diff_within_at 𝕜 s x) - (hL : f₁ =ᶠ[𝓝[s] x] f) (hx : f₁ x = f x) : +lemma filter.eventually_eq.deriv_within_eq (hL : f₁ =ᶠ[𝓝[s] x] f) (hx : f₁ x = f x) : deriv_within f₁ s x = deriv_within f s x := -by { unfold deriv_within, rw hL.fderiv_within_eq hs hx } +by { unfold deriv_within, rw hL.fderiv_within_eq hx } -lemma deriv_within_congr (hs : unique_diff_within_at 𝕜 s x) - (hL : ∀y∈s, f₁ y = f y) (hx : f₁ x = f x) : +lemma deriv_within_congr (hs : eq_on f₁ f s) (hx : f₁ x = f x) : deriv_within f₁ s x = deriv_within f s x := -by { unfold deriv_within, rw fderiv_within_congr hs hL hx } +by { unfold deriv_within, rw fderiv_within_congr hs hx } lemma filter.eventually_eq.deriv_eq (hL : f₁ =ᶠ[𝓝 x] f) : deriv f₁ x = deriv f x := by { unfold deriv, rwa filter.eventually_eq.fderiv_eq } diff --git a/src/analysis/calculus/fderiv/basic.lean b/src/analysis/calculus/fderiv/basic.lean index 566a26ac49a15..b7fcebc366f5d 100644 --- a/src/analysis/calculus/fderiv/basic.lean +++ b/src/analysis/calculus/fderiv/basic.lean @@ -352,23 +352,27 @@ by { simp only [has_fderiv_within_at, nhds_within_univ], refl } alias has_fderiv_within_at_univ ↔ has_fderiv_within_at.has_fderiv_at_of_univ _ -lemma has_fderiv_within_at_insert {y : E} {g' : E →L[𝕜] F} : - has_fderiv_within_at g g' (insert y s) x ↔ has_fderiv_within_at g g' s x := +lemma has_fderiv_within_at_insert {y : E} : + has_fderiv_within_at f f' (insert y s) x ↔ has_fderiv_within_at f f' s x := begin rcases eq_or_ne x y with rfl|h, { simp_rw [has_fderiv_within_at, has_fderiv_at_filter], apply asymptotics.is_o_insert, - simp only [sub_self, g'.map_zero] }, - refine ⟨λ h, h.mono $ subset_insert y s, λ hg, hg.mono_of_mem _⟩, + simp only [sub_self, map_zero] }, + refine ⟨λ h, h.mono $ subset_insert y s, λ hf, hf.mono_of_mem _⟩, simp_rw [nhds_within_insert_of_ne h, self_mem_nhds_within] end alias has_fderiv_within_at_insert ↔ has_fderiv_within_at.of_insert has_fderiv_within_at.insert' -lemma has_fderiv_within_at.insert {g' : E →L[𝕜] F} (h : has_fderiv_within_at g g' s x) : - has_fderiv_within_at g g' (insert x s) x := +lemma has_fderiv_within_at.insert (h : has_fderiv_within_at f f' s x) : + has_fderiv_within_at f f' (insert x s) x := h.insert' +lemma has_fderiv_within_at_diff_singleton (y : E) : + has_fderiv_within_at f f' (s \ {y}) x ↔ has_fderiv_within_at f f' s x := +by rw [← has_fderiv_within_at_insert, insert_diff_singleton, has_fderiv_within_at_insert] + lemma has_strict_fderiv_at.is_O_sub (hf : has_strict_fderiv_at f f' x) : (λ p : E × E, f p.1 - f p.2) =O[𝓝 (x, x)] (λ p : E × E, p.1 - p.2) := hf.is_O.congr_of_sub.2 (f'.is_O_comp _ _) @@ -520,7 +524,7 @@ begin end lemma differentiable_within_at.mono_of_mem (h : differentiable_within_at 𝕜 f s x) {t : set E} - (hst : s ∈ nhds_within x t) : + (hst : s ∈ 𝓝[t] x) : differentiable_within_at 𝕜 f t x := (h.has_fderiv_within_at.mono_of_mem hst).differentiable_within_at @@ -530,27 +534,11 @@ by simp only [differentiable_within_at, has_fderiv_within_at_univ, differentiabl lemma differentiable_within_at_inter (ht : t ∈ 𝓝 x) : differentiable_within_at 𝕜 f (s ∩ t) x ↔ differentiable_within_at 𝕜 f s x := -by simp only [differentiable_within_at, has_fderiv_within_at, has_fderiv_at_filter, - nhds_within_restrict' s ht] +by simp only [differentiable_within_at, has_fderiv_within_at_inter ht] lemma differentiable_within_at_inter' (ht : t ∈ 𝓝[s] x) : differentiable_within_at 𝕜 f (s ∩ t) x ↔ differentiable_within_at 𝕜 f s x := -by simp only [differentiable_within_at, has_fderiv_within_at, has_fderiv_at_filter, - nhds_within_restrict'' s ht] - -lemma differentiable_within_at.antimono (h : differentiable_within_at 𝕜 f s x) (hst : s ⊆ t) - (hx : s ∈ 𝓝[t] x) : - differentiable_within_at 𝕜 f t x := -by rwa [← differentiable_within_at_inter' hx, inter_eq_self_of_subset_right hst] - -lemma has_fderiv_within_at.antimono (h : has_fderiv_within_at f f' s x) (hst : s ⊆ t) - (hs : unique_diff_within_at 𝕜 s x) (hx : s ∈ 𝓝[t] x) : - has_fderiv_within_at f f' t x := -begin - have h' : has_fderiv_within_at f _ t x := - (h.differentiable_within_at.antimono hst hx).has_fderiv_within_at, - rwa hs.eq h (h'.mono hst), -end +by simp only [differentiable_within_at, has_fderiv_within_at_inter' ht] lemma differentiable_at.differentiable_within_at (h : differentiable_at 𝕜 f x) : differentiable_within_at 𝕜 f s x := @@ -585,52 +573,30 @@ begin exact (differentiable_within_at_inter (is_open.mem_nhds t_open xt)).1 (ht x ⟨xs, xt⟩) end -lemma fderiv_within_subset (st : s ⊆ t) (ht : unique_diff_within_at 𝕜 s x) +lemma fderiv_within_of_mem (st : t ∈ 𝓝[s] x) (ht : unique_diff_within_at 𝕜 s x) (h : differentiable_within_at 𝕜 f t x) : fderiv_within 𝕜 f s x = fderiv_within 𝕜 f t x := -((differentiable_within_at.has_fderiv_within_at h).mono st).fderiv_within ht +((differentiable_within_at.has_fderiv_within_at h).mono_of_mem st).fderiv_within ht -lemma fderiv_within_subset' (st : s ⊆ t) (ht : unique_diff_within_at 𝕜 s x) (hx : s ∈ 𝓝[t] x) - (h : differentiable_within_at 𝕜 f s x) : +lemma fderiv_within_subset (st : s ⊆ t) (ht : unique_diff_within_at 𝕜 s x) + (h : differentiable_within_at 𝕜 f t x) : fderiv_within 𝕜 f s x = fderiv_within 𝕜 f t x := -fderiv_within_subset st ht (h.antimono st hx) +fderiv_within_of_mem (nhds_within_mono _ st self_mem_nhds_within) ht h -@[simp] lemma fderiv_within_univ : fderiv_within 𝕜 f univ = fderiv 𝕜 f := -begin - ext x : 1, - by_cases h : differentiable_at 𝕜 f x, - { apply has_fderiv_within_at.fderiv_within _ unique_diff_within_at_univ, - rw has_fderiv_within_at_univ, - apply h.has_fderiv_at }, - { have : ¬ differentiable_within_at 𝕜 f univ x, - { rwa differentiable_within_at_univ }, - rw [fderiv_zero_of_not_differentiable_at h, - fderiv_within_zero_of_not_differentiable_within_at this] } -end - -lemma fderiv_within_inter (ht : t ∈ 𝓝 x) (hs : unique_diff_within_at 𝕜 s x) : +lemma fderiv_within_inter (ht : t ∈ 𝓝 x) : fderiv_within 𝕜 f (s ∩ t) x = fderiv_within 𝕜 f s x := -begin - by_cases h : differentiable_within_at 𝕜 f (s ∩ t) x, - { apply fderiv_within_subset (inter_subset_left _ _) _ ((differentiable_within_at_inter ht).1 h), - apply hs.inter ht }, - { have : ¬ differentiable_within_at 𝕜 f s x, - { rwa ←differentiable_within_at_inter ht }, - rw [fderiv_within_zero_of_not_differentiable_within_at h, - fderiv_within_zero_of_not_differentiable_within_at this] } -end +by simp only [fderiv_within, has_fderiv_within_at_inter ht] lemma fderiv_within_of_mem_nhds (h : s ∈ 𝓝 x) : fderiv_within 𝕜 f s x = fderiv 𝕜 f x := -begin - have : s = univ ∩ s, by simp only [univ_inter], - rw [this, ← fderiv_within_univ], - exact fderiv_within_inter h (unique_diff_on_univ _ (mem_univ _)) -end +by simp only [fderiv, fderiv_within, has_fderiv_at, has_fderiv_within_at, nhds_within_eq_nhds.2 h] + +@[simp] lemma fderiv_within_univ : fderiv_within 𝕜 f univ = fderiv 𝕜 f := +funext $ λ _, fderiv_within_of_mem_nhds univ_mem lemma fderiv_within_of_open (hs : is_open s) (hx : x ∈ s) : fderiv_within 𝕜 f s x = fderiv 𝕜 f x := -fderiv_within_of_mem_nhds (is_open.mem_nhds hs hx) +fderiv_within_of_mem_nhds (hs.mem_nhds hx) lemma fderiv_within_eq_fderiv (hs : unique_diff_within_at 𝕜 s x) (h : differentiable_at 𝕜 f x) : fderiv_within 𝕜 f s x = fderiv 𝕜 f x := @@ -735,6 +701,44 @@ end continuous section congr /-! ### congr properties of the derivative -/ +lemma has_fderiv_within_at_congr_set' (y : E) (h : s =ᶠ[𝓝[{y}ᶜ] x] t) : + has_fderiv_within_at f f' s x ↔ has_fderiv_within_at f f' t x := +calc has_fderiv_within_at f f' s x ↔ has_fderiv_within_at f f' (s \ {y}) x : + (has_fderiv_within_at_diff_singleton _).symm +... ↔ has_fderiv_within_at f f' (t \ {y}) x : + suffices 𝓝[s \ {y}] x = 𝓝[t \ {y}] x, by simp only [has_fderiv_within_at, this], + by simpa only [set_eventually_eq_iff_inf_principal, ← nhds_within_inter', diff_eq, inter_comm] + using h +... ↔ has_fderiv_within_at f f' t x : has_fderiv_within_at_diff_singleton _ + +lemma has_fderiv_within_at_congr_set (h : s =ᶠ[𝓝 x] t) : + has_fderiv_within_at f f' s x ↔ has_fderiv_within_at f f' t x := +has_fderiv_within_at_congr_set' x $ h.filter_mono inf_le_left + +lemma differentiable_within_at_congr_set' (y : E) (h : s =ᶠ[𝓝[{y}ᶜ] x] t) : + differentiable_within_at 𝕜 f s x ↔ differentiable_within_at 𝕜 f t x := +exists_congr $ λ _, has_fderiv_within_at_congr_set' _ h + +lemma differentiable_within_at_congr_set (h : s =ᶠ[𝓝 x] t) : + differentiable_within_at 𝕜 f s x ↔ differentiable_within_at 𝕜 f t x := +exists_congr $ λ _, has_fderiv_within_at_congr_set h + +lemma fderiv_within_congr_set' (y : E) (h : s =ᶠ[𝓝[{y}ᶜ] x] t) : + fderiv_within 𝕜 f s x = fderiv_within 𝕜 f t x := +by simp only [fderiv_within, has_fderiv_within_at_congr_set' y h] + +lemma fderiv_within_congr_set (h : s =ᶠ[𝓝 x] t) : + fderiv_within 𝕜 f s x = fderiv_within 𝕜 f t x := +fderiv_within_congr_set' x $ h.filter_mono inf_le_left + +lemma fderiv_within_eventually_congr_set' (y : E) (h : s =ᶠ[𝓝[{y}ᶜ] x] t) : + fderiv_within 𝕜 f s =ᶠ[𝓝 x] fderiv_within 𝕜 f t := +(eventually_nhds_nhds_within.2 h).mono $ λ _, fderiv_within_congr_set' y + +lemma fderiv_within_eventually_congr_set (h : s =ᶠ[𝓝 x] t) : + fderiv_within 𝕜 f s =ᶠ[𝓝 x] fderiv_within 𝕜 f t := +fderiv_within_eventually_congr_set' x $ h.filter_mono inf_le_left + theorem filter.eventually_eq.has_strict_fderiv_at_iff (h : f₀ =ᶠ[𝓝 x] f₁) (h' : ∀ y, f₀' y = f₁' y) : has_strict_fderiv_at f₀ f₀' x ↔ has_strict_fderiv_at f₁ f₁' x := @@ -783,17 +787,17 @@ theorem filter.eventually_eq.differentiable_within_at_iff_of_mem (h : f₀ =ᶠ[ differentiable_within_at 𝕜 f₀ s x ↔ differentiable_within_at 𝕜 f₁ s x := h.differentiable_within_at_iff (h.eq_of_nhds_within hx) -lemma has_fderiv_within_at.congr_mono (h : has_fderiv_within_at f f' s x) (ht : ∀x ∈ t, f₁ x = f x) +lemma has_fderiv_within_at.congr_mono (h : has_fderiv_within_at f f' s x) (ht : eq_on f₁ f t) (hx : f₁ x = f x) (h₁ : t ⊆ s) : has_fderiv_within_at f₁ f' t x := has_fderiv_at_filter.congr_of_eventually_eq (h.mono h₁) (filter.mem_inf_of_right ht) hx -lemma has_fderiv_within_at.congr (h : has_fderiv_within_at f f' s x) (hs : ∀x ∈ s, f₁ x = f x) +lemma has_fderiv_within_at.congr (h : has_fderiv_within_at f f' s x) (hs : eq_on f₁ f s) (hx : f₁ x = f x) : has_fderiv_within_at f₁ f' s x := h.congr_mono hs hx (subset.refl _) -lemma has_fderiv_within_at.congr' (h : has_fderiv_within_at f f' s x) (hs : ∀x ∈ s, f₁ x = f x) +lemma has_fderiv_within_at.congr' (h : has_fderiv_within_at f f' s x) (hs : eq_on f₁ f s) (hx : x ∈ s) : has_fderiv_within_at f₁ f' s x := -h.congr hs (hs x hx) +h.congr hs (hs hx) lemma has_fderiv_within_at.congr_of_eventually_eq (h : has_fderiv_within_at f f' s x) (h₁ : f₁ =ᶠ[𝓝[s] x] f) (hx : f₁ x = f x) : has_fderiv_within_at f₁ f' s x := @@ -804,8 +808,8 @@ lemma has_fderiv_at.congr_of_eventually_eq (h : has_fderiv_at f f' x) has_fderiv_at_filter.congr_of_eventually_eq h h₁ (mem_of_mem_nhds h₁ : _) lemma differentiable_within_at.congr_mono (h : differentiable_within_at 𝕜 f s x) - (ht : ∀x ∈ t, f₁ x = f x) (hx : f₁ x = f x) (h₁ : t ⊆ s) : differentiable_within_at 𝕜 f₁ t x := -(has_fderiv_within_at.congr_mono h.has_fderiv_within_at ht hx h₁).differentiable_within_at + (ht : eq_on f₁ f t) (hx : f₁ x = f x) (h₁ : t ⊆ s) : differentiable_within_at 𝕜 f₁ t x := +(h.has_fderiv_within_at.congr_mono ht hx h₁).differentiable_within_at lemma differentiable_within_at.congr (h : differentiable_within_at 𝕜 f s x) (ht : ∀x ∈ s, f₁ x = f x) (hx : f₁ x = f x) : differentiable_within_at 𝕜 f₁ s x := @@ -834,48 +838,39 @@ lemma differentiable_at.congr_of_eventually_eq (h : differentiable_at 𝕜 f x) hL.differentiable_at_iff.2 h lemma differentiable_within_at.fderiv_within_congr_mono (h : differentiable_within_at 𝕜 f s x) - (hs : ∀x ∈ t, f₁ x = f x) (hx : f₁ x = f x) (hxt : unique_diff_within_at 𝕜 t x) (h₁ : t ⊆ s) : + (hs : eq_on f₁ f t) (hx : f₁ x = f x) (hxt : unique_diff_within_at 𝕜 t x) (h₁ : t ⊆ s) : fderiv_within 𝕜 f₁ t x = fderiv_within 𝕜 f s x := (has_fderiv_within_at.congr_mono h.has_fderiv_within_at hs hx h₁).fderiv_within hxt -lemma filter.eventually_eq.fderiv_within_eq (hs : unique_diff_within_at 𝕜 s x) - (hL : f₁ =ᶠ[𝓝[s] x] f) (hx : f₁ x = f x) : +lemma filter.eventually_eq.fderiv_within_eq (hs : f₁ =ᶠ[𝓝[s] x] f) (hx : f₁ x = f x) : fderiv_within 𝕜 f₁ s x = fderiv_within 𝕜 f s x := -if h : differentiable_within_at 𝕜 f s x -then has_fderiv_within_at.fderiv_within (h.has_fderiv_within_at.congr_of_eventually_eq hL hx) hs -else - have h' : ¬ differentiable_within_at 𝕜 f₁ s x, - from mt (λ h, h.congr_of_eventually_eq (hL.mono $ λ x, eq.symm) hx.symm) h, - by rw [fderiv_within_zero_of_not_differentiable_within_at h, - fderiv_within_zero_of_not_differentiable_within_at h'] - -lemma filter.eventually_eq.fderiv_within_eq_nhds (hs : unique_diff_within_at 𝕜 s x) - (hL : f₁ =ᶠ[𝓝 x] f) : +by simp only [fderiv_within, hs.has_fderiv_within_at_iff hx] + +lemma filter.eventually_eq.fderiv_within' (hs : f₁ =ᶠ[𝓝[s] x] f) (ht : t ⊆ s) : + fderiv_within 𝕜 f₁ t =ᶠ[𝓝[s] x] fderiv_within 𝕜 f t := +(eventually_nhds_within_nhds_within.2 hs).mp $ eventually_mem_nhds_within.mono $ λ y hys hs, + filter.eventually_eq.fderiv_within_eq (hs.filter_mono $ nhds_within_mono _ ht) + (hs.self_of_nhds_within hys) + +protected lemma filter.eventually_eq.fderiv_within (hs : f₁ =ᶠ[𝓝[s] x] f) : + fderiv_within 𝕜 f₁ s =ᶠ[𝓝[s] x] fderiv_within 𝕜 f s := +hs.fderiv_within' subset.rfl + +lemma filter.eventually_eq.fderiv_within_eq_nhds (h : f₁ =ᶠ[𝓝 x] f) : fderiv_within 𝕜 f₁ s x = fderiv_within 𝕜 f s x := -(show f₁ =ᶠ[𝓝[s] x] f, from nhds_within_le_nhds hL).fderiv_within_eq hs (mem_of_mem_nhds hL : _) +(h.filter_mono nhds_within_le_nhds).fderiv_within_eq h.self_of_nhds -lemma fderiv_within_congr (hs : unique_diff_within_at 𝕜 s x) - (hL : ∀ y ∈ s, f₁ y = f y) (hx : f₁ x = f x) : +lemma fderiv_within_congr (hs : eq_on f₁ f s) (hx : f₁ x = f x) : fderiv_within 𝕜 f₁ s x = fderiv_within 𝕜 f s x := -begin - apply filter.eventually_eq.fderiv_within_eq hs _ hx, - apply mem_of_superset self_mem_nhds_within, - exact hL -end +(hs.eventually_eq.filter_mono inf_le_right).fderiv_within_eq hx -lemma fderiv_within_congr' (hs : unique_diff_within_at 𝕜 s x) - (hL : ∀ y ∈ s, f₁ y = f y) (hx : x ∈ s) : +lemma fderiv_within_congr' (hs : eq_on f₁ f s) (hx : x ∈ s) : fderiv_within 𝕜 f₁ s x = fderiv_within 𝕜 f s x := -fderiv_within_congr hs hL (hL x hx) +fderiv_within_congr hs (hs hx) -lemma filter.eventually_eq.fderiv_eq (hL : f₁ =ᶠ[𝓝 x] f) : +lemma filter.eventually_eq.fderiv_eq (h : f₁ =ᶠ[𝓝 x] f) : fderiv 𝕜 f₁ x = fderiv 𝕜 f x := -begin - have A : f₁ x = f x := hL.eq_of_nhds, - rw [← fderiv_within_univ, ← fderiv_within_univ], - rw ← nhds_within_univ at hL, - exact hL.fderiv_within_eq unique_diff_within_at_univ A -end +by rw [← fderiv_within_univ, ← fderiv_within_univ, h.fderiv_within_eq_nhds] protected lemma filter.eventually_eq.fderiv (h : f₁ =ᶠ[𝓝 x] f) : fderiv 𝕜 f₁ =ᶠ[𝓝 x] fderiv 𝕜 f := @@ -1022,12 +1017,13 @@ variables (𝕜 : Type*) {E F : Type*} [nontrivially_normed_field 𝕜] [normed_ lemma support_fderiv_subset : support (fderiv 𝕜 f) ⊆ tsupport f := begin intros x, - rw [← not_imp_not], - intro h2x, - rw [not_mem_tsupport_iff_eventually_eq] at h2x, - exact nmem_support.mpr (h2x.fderiv_eq.trans $ fderiv_const_apply 0), + rw [← not_imp_not, not_mem_tsupport_iff_eventually_eq, nmem_support], + exact λ hx, (hx.fderiv_eq.trans $ fderiv_const_apply 0), end +lemma tsupport_fderiv_subset : tsupport (fderiv 𝕜 f) ⊆ tsupport f := +closure_minimal (support_fderiv_subset 𝕜) is_closed_closure + lemma has_compact_support.fderiv (hf : has_compact_support f) : has_compact_support (fderiv 𝕜 f) := hf.mono' $ support_fderiv_subset 𝕜 diff --git a/src/analysis/calculus/iterated_deriv.lean b/src/analysis/calculus/iterated_deriv.lean index 67837508500c8..58d836ff92a25 100644 --- a/src/analysis/calculus/iterated_deriv.lean +++ b/src/analysis/calculus/iterated_deriv.lean @@ -104,9 +104,9 @@ by rw [iterated_deriv_within_eq_equiv_comp, linear_isometry_equiv.norm_map] iterated_deriv_within 0 f s = f := by { ext x, simp [iterated_deriv_within] } -@[simp] lemma iterated_deriv_within_one (hs : unique_diff_on 𝕜 s) {x : 𝕜} (hx : x ∈ s): +@[simp] lemma iterated_deriv_within_one {x : 𝕜} (h : unique_diff_within_at 𝕜 s x): iterated_deriv_within 1 f s x = deriv_within f s x := -by { simp [iterated_deriv_within, iterated_fderiv_within_one_apply hs hx], refl } +by { simp only [iterated_deriv_within, iterated_fderiv_within_one_apply h], refl } /-- If the first `n` derivatives within a set of a function are continuous, and its first `n-1` derivatives are differentiable, then the function is `C^n`. This is not an equivalence in general, @@ -146,14 +146,19 @@ lemma cont_diff_on.continuous_on_iterated_deriv_within {n : ℕ∞} {m : ℕ} by simpa only [iterated_deriv_within_eq_equiv_comp, linear_isometry_equiv.comp_continuous_on_iff] using h.continuous_on_iterated_fderiv_within hmn hs +lemma cont_diff_within_at.differentiable_within_at_iterated_deriv_within {n : ℕ∞} {m : ℕ} + (h : cont_diff_within_at 𝕜 n f s x) (hmn : (m : ℕ∞) < n) (hs : unique_diff_on 𝕜 (insert x s)) : + differentiable_within_at 𝕜 (iterated_deriv_within m f s) s x := +by simpa only [iterated_deriv_within_eq_equiv_comp, + linear_isometry_equiv.comp_differentiable_within_at_iff] + using h.differentiable_within_at_iterated_fderiv_within hmn hs + /-- On a set with unique derivatives, a `C^n` function has derivatives less than `n` which are differentiable. -/ lemma cont_diff_on.differentiable_on_iterated_deriv_within {n : ℕ∞} {m : ℕ} (h : cont_diff_on 𝕜 n f s) (hmn : (m : ℕ∞) < n) (hs : unique_diff_on 𝕜 s) : differentiable_on 𝕜 (iterated_deriv_within m f s) s := -by simpa only [iterated_deriv_within_eq_equiv_comp, - linear_isometry_equiv.comp_differentiable_on_iff] - using h.differentiable_on_iterated_fderiv_within hmn hs +λ x hx, (h x hx).differentiable_within_at_iterated_deriv_within hmn $ by rwa [insert_eq_of_mem hx] /-- The property of being `C^n`, initially defined in terms of the Fréchet derivative, can be reformulated in terms of the one-dimensional derivative on sets with unique derivatives. -/ @@ -189,7 +194,7 @@ begin induction n with n IH generalizing x, { simp }, { rw [iterated_deriv_within_succ (hs x hx), function.iterate_succ'], - exact deriv_within_congr (hs x hx) (λ y hy, IH hy) (IH hx) } + exact deriv_within_congr (λ y hy, IH hy) (IH hx) } end /-- The `n+1`-th iterated derivative within a set with unique derivatives can be obtained by diff --git a/src/analysis/calculus/taylor.lean b/src/analysis/calculus/taylor.lean index 0ad33d0d28309..eb40d84acfff6 100644 --- a/src/analysis/calculus/taylor.lean +++ b/src/analysis/calculus/taylor.lean @@ -146,34 +146,29 @@ begin simp only [nat.cast_add, nat.cast_one], end -lemma has_deriv_within_at_taylor_coeff_within {f : ℝ → E} {x y : ℝ} {k : ℕ} {s s' : set ℝ} - (hs'_unique : unique_diff_within_at ℝ s' y) - (hs' : s' ∈ 𝓝[s] y) (hy : y ∈ s') (h : s' ⊆ s) - (hf' : differentiable_on ℝ (iterated_deriv_within (k+1) f s) s') : - has_deriv_within_at (λ t, - (((k+1 : ℝ) * k!)⁻¹ * (x - t)^(k+1)) • iterated_deriv_within (k+1) f s t) +lemma has_deriv_within_at_taylor_coeff_within {f : ℝ → E} {x y : ℝ} {k : ℕ} {s t : set ℝ} + (ht : unique_diff_within_at ℝ t y) (hs : s ∈ 𝓝[t] y) + (hf : differentiable_within_at ℝ (iterated_deriv_within (k+1) f s) s y) : + has_deriv_within_at (λ z, + (((k+1 : ℝ) * k!)⁻¹ * (x - z)^(k+1)) • iterated_deriv_within (k+1) f s z) ((((k+1 : ℝ) * k!)⁻¹ * (x - y)^(k+1)) • iterated_deriv_within (k+2) f s y - - ((k! : ℝ)⁻¹ * (x - y)^k) • iterated_deriv_within (k+1) f s y) s' y := + ((k! : ℝ)⁻¹ * (x - y)^k) • iterated_deriv_within (k+1) f s y) t y := begin - have hf'' : has_deriv_within_at (λ t, iterated_deriv_within (k+1) f s t) - (iterated_deriv_within (k+2) f s y) s' y := + replace hf : has_deriv_within_at (iterated_deriv_within (k+1) f s) + (iterated_deriv_within (k+2) f s y) t y := begin - convert (hf' y hy).has_deriv_within_at, - rw iterated_deriv_within_succ (hs'_unique.mono h), - refine (deriv_within_subset h hs'_unique _).symm, - exact (hf' y hy).antimono h hs', + convert (hf.mono_of_mem hs).has_deriv_within_at, + rw iterated_deriv_within_succ (ht.mono_nhds (nhds_within_le_iff.mpr hs)), + exact (deriv_within_of_mem hs ht hf).symm end, have : has_deriv_within_at (λ t, (((k+1 : ℝ) * k!)⁻¹ * (x - t)^(k+1))) - (-((k! : ℝ)⁻¹ * (x - y)^k)) s' y := - begin - -- Commuting the factors: - have : (-((k! : ℝ)⁻¹ * (x - y)^k)) = - (((k+1 : ℝ) * k!)⁻¹ * (-(k+1) *(x - y)^k)) := - by { field_simp [nat.cast_add_one_ne_zero k, nat.factorial_ne_zero k], ring_nf }, + (-((k! : ℝ)⁻¹ * (x - y)^k)) t y, + { -- Commuting the factors: + have : (-((k! : ℝ)⁻¹ * (x - y)^k)) = (((k+1 : ℝ) * k!)⁻¹ * (-(k+1) *(x - y)^k)), + { field_simp [nat.cast_add_one_ne_zero k, nat.factorial_ne_zero k], ring_nf }, rw this, - exact (monomial_has_deriv_aux y x _).has_deriv_within_at.const_mul _, - end, - convert this.smul hf'', + exact (monomial_has_deriv_aux y x _).has_deriv_within_at.const_mul _ }, + convert this.smul hf, field_simp [nat.cast_add_one_ne_zero k, nat.factorial_ne_zero k], rw [neg_div, neg_smul, sub_eq_add_neg], end @@ -185,7 +180,7 @@ lemma has_deriv_within_at_taylor_within_eval {f : ℝ → E} {x y : ℝ} {n : (hs'_unique : unique_diff_within_at ℝ s' y) (hs_unique : unique_diff_on ℝ s) (hs' : s' ∈ 𝓝[s] y) (hy : y ∈ s') (h : s' ⊆ s) (hf : cont_diff_on ℝ n f s) - (hf' : differentiable_on ℝ (iterated_deriv_within n f s) s') : + (hf' : differentiable_within_at ℝ (iterated_deriv_within n f s) s y) : has_deriv_within_at (λ t, taylor_within_eval f n s t x) (((n! : ℝ)⁻¹ * (x - y)^n) • (iterated_deriv_within (n+1) f s y)) s' y := begin @@ -193,22 +188,18 @@ begin { simp only [taylor_within_zero_eval, nat.factorial_zero, nat.cast_one, inv_one, pow_zero, mul_one, zero_add, one_smul], simp only [iterated_deriv_within_zero] at hf', - rw iterated_deriv_within_one hs_unique (h hy), - refine has_deriv_within_at.mono _ h, - refine differentiable_within_at.has_deriv_within_at _, - exact (hf' y hy).antimono h hs' }, + rw iterated_deriv_within_one (hs_unique _ (h hy)), + exact hf'.has_deriv_within_at.mono h }, simp_rw [nat.add_succ, taylor_within_eval_succ], simp only [add_zero, nat.factorial_succ, nat.cast_mul, nat.cast_add, nat.cast_one], - have hdiff : differentiable_on ℝ (iterated_deriv_within k f s) s' := - begin - have coe_lt_succ : (k : with_top ℕ) < k.succ := - by { rw [with_top.coe_lt_coe], exact lt_add_one k }, + have hdiff : differentiable_on ℝ (iterated_deriv_within k f s) s', + { have coe_lt_succ : (k : with_top ℕ) < k.succ := nat.cast_lt.2 k.lt_succ_self, refine differentiable_on.mono _ h, - exact hf.differentiable_on_iterated_deriv_within coe_lt_succ hs_unique, - end, - specialize hk (cont_diff_on.of_succ hf) hdiff, - convert hk.add (has_deriv_within_at_taylor_coeff_within hs'_unique hs' hy h hf'), - exact (add_sub_cancel'_right _ _).symm, + exact hf.differentiable_on_iterated_deriv_within coe_lt_succ hs_unique }, + specialize hk hf.of_succ ((hdiff y hy).mono_of_mem hs'), + convert hk.add (has_deriv_within_at_taylor_coeff_within hs'_unique + (nhds_within_mono _ h self_mem_nhds_within) hf'), + exact (add_sub_cancel'_right _ _).symm end /-- Calculate the derivative of the Taylor polynomial with respect to `x₀`. @@ -220,12 +211,10 @@ lemma taylor_within_eval_has_deriv_at_Ioo {f : ℝ → E} {a b t : ℝ} (x : ℝ (hf' : differentiable_on ℝ (iterated_deriv_within n f (Icc a b)) (Ioo a b)) : has_deriv_at (λ y, taylor_within_eval f n (Icc a b) y x) (((n! : ℝ)⁻¹ * (x - t)^n) • (iterated_deriv_within (n+1) f (Icc a b) t)) t := -begin - have h_nhds := is_open.mem_nhds is_open_Ioo ht, - exact (has_deriv_within_at_taylor_within_eval (unique_diff_within_at_Ioo ht) - (unique_diff_on_Icc hx) (nhds_within_le_nhds h_nhds) ht Ioo_subset_Icc_self hf hf') - .has_deriv_at h_nhds, -end +have h_nhds : Ioo a b ∈ 𝓝 t := is_open_Ioo.mem_nhds ht, +have h_nhds' : Ioo a b ∈ 𝓝[Icc a b] t := nhds_within_le_nhds h_nhds, +(has_deriv_within_at_taylor_within_eval (unique_diff_within_at_Ioo ht) (unique_diff_on_Icc hx) + h_nhds' ht Ioo_subset_Icc_self hf $ (hf' t ht).mono_of_mem h_nhds').has_deriv_at h_nhds /-- Calculate the derivative of the Taylor polynomial with respect to `x₀`. @@ -236,7 +225,7 @@ lemma has_deriv_within_taylor_within_eval_at_Icc {f : ℝ → E} {a b t : ℝ} ( has_deriv_within_at (λ y, taylor_within_eval f n (Icc a b) y x) (((n! : ℝ)⁻¹ * (x - t)^n) • (iterated_deriv_within (n+1) f (Icc a b) t)) (Icc a b) t := has_deriv_within_at_taylor_within_eval (unique_diff_on_Icc hx t ht) (unique_diff_on_Icc hx) - self_mem_nhds_within ht rfl.subset hf hf' + self_mem_nhds_within ht rfl.subset hf (hf' t ht) /-! ### Taylor's theorem with mean value type remainder estimate -/ diff --git a/src/geometry/manifold/cont_mdiff_mfderiv.lean b/src/geometry/manifold/cont_mdiff_mfderiv.lean index 3da2e58fd45f5..343ec2c230329 100644 --- a/src/geometry/manifold/cont_mdiff_mfderiv.lean +++ b/src/geometry/manifold/cont_mdiff_mfderiv.lean @@ -158,7 +158,7 @@ begin { rintro x ⟨hx, h2x⟩, simp_rw [written_in_ext_chart_at, function.comp_apply], rw [(ext_chart_at I (g x₂)).left_inv hx, (ext_chart_at I' (f x₂ (g x₂))).left_inv h2x] }, - refine filter.eventually_eq.fderiv_within_eq_nhds (I.unique_diff _ $ mem_range_self _) _, + refine filter.eventually_eq.fderiv_within_eq_nhds _, refine eventually_of_mem (inter_mem _ _) this, { exact ext_chart_at_preimage_mem_nhds' _ _ hx₂ (ext_chart_at_source_mem_nhds I (g x₂)) }, refine ext_chart_at_preimage_mem_nhds' _ _ hx₂ _, @@ -599,15 +599,13 @@ begin simp only [bundle.zero_section, tangent_map, mfderiv, total_space.proj_mk, A, if_pos, chart_at, fiber_bundle.charted_space_chart_at, tangent_bundle.trivialization_at_apply, tangent_bundle_core, function.comp, continuous_linear_map.map_zero] with mfld_simps, - rw ← fderiv_within_inter N (I.unique_diff (I ((chart_at H x) x)) (set.mem_range_self _)) at B, - rw [← fderiv_within_inter N (I.unique_diff (I ((chart_at H x) x)) (set.mem_range_self _)), ← B], + rw [← fderiv_within_inter N] at B, + rw [← fderiv_within_inter N, ← B], congr' 2, - apply fderiv_within_congr _ (λ y hy, _), - { simp only [prod.mk.inj_iff] with mfld_simps }, - { apply unique_diff_within_at.inter (I.unique_diff _ _) N, - simp only with mfld_simps }, + refine fderiv_within_congr (λ y hy, _) _, { simp only with mfld_simps at hy, simp only [hy, prod.mk.inj_iff] with mfld_simps }, + { simp only [prod.mk.inj_iff] with mfld_simps }, end end tangent_bundle diff --git a/src/geometry/manifold/mfderiv.lean b/src/geometry/manifold/mfderiv.lean index e80c4ec556fb4..971283f6f685a 100644 --- a/src/geometry/manifold/mfderiv.lean +++ b/src/geometry/manifold/mfderiv.lean @@ -616,11 +616,11 @@ begin rw mdifferentiable_within_at_univ end -lemma mfderiv_within_inter (ht : t ∈ 𝓝 x) (hs : unique_mdiff_within_at I s x) : +lemma mfderiv_within_inter (ht : t ∈ 𝓝 x) : mfderiv_within I I' f (s ∩ t) x = mfderiv_within I I' f s x := by rw [mfderiv_within, mfderiv_within, ext_chart_at_preimage_inter_eq, mdifferentiable_within_at_inter ht, - fderiv_within_inter (ext_chart_at_preimage_mem_nhds I x ht) hs] + fderiv_within_inter (ext_chart_at_preimage_mem_nhds I x ht)] lemma mdifferentiable_at_iff_of_mem_source {x' : M} {y : M'} (hx : x' ∈ (charted_space.chart_at H x).source) diff --git a/src/geometry/manifold/vector_bundle/tangent.lean b/src/geometry/manifold/vector_bundle/tangent.lean index bd35638c65230..f8942eb407f51 100644 --- a/src/geometry/manifold/vector_bundle/tangent.lean +++ b/src/geometry/manifold/vector_bundle/tangent.lean @@ -83,7 +83,6 @@ within the set `range I ⊆ E` at `I (i x) : E`. -/ coord_change_self := λ i x hx v, begin rw [filter.eventually_eq.fderiv_within_eq, fderiv_within_id', continuous_linear_map.id_apply], { exact I.unique_diff_at_image }, - { exact I.unique_diff_at_image }, { filter_upwards [i.1.extend_target_mem_nhds_within I hx] with y hy, exact (i.1.extend I).right_inv hy }, { simp_rw [function.comp_apply, i.1.extend_left_inv I hx] } @@ -97,7 +96,6 @@ within the set `range I ⊆ E` at `I (i x) : E`. -/ coord_change_comp := begin rintro i j k x ⟨⟨hxi, hxj⟩, hxk⟩ v, rw [fderiv_within_fderiv_within, filter.eventually_eq.fderiv_within_eq], - { exact I.unique_diff_at_image }, { have := i.1.extend_preimage_mem_nhds I hxi (j.1.extend_source_mem_nhds I hxj), filter_upwards [nhds_within_le_nhds this] with y hy, simp_rw [function.comp_apply, (j.1.extend I).left_inv hy] }, From 61db041ab8e4aaf8cb5c7dc10a7d4ff261997536 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Wed, 24 May 2023 04:13:02 +0000 Subject: [PATCH 1067/1291] chore(*): add mathlib4 synchronization comments (#19076) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.category.Ring.colimits` * `analysis.box_integral.partition.additive` * `analysis.calculus.formal_multilinear_series` * `analysis.convex.cone.basic` * `analysis.convex.side` * `data.nat.prime_norm_num` * `field_theory.ax_grothendieck` * `field_theory.intermediate_field` * `field_theory.separable` * `linear_algebra.annihilating_polynomial` * `linear_algebra.charpoly.basic` * `linear_algebra.free_module.strong_rank_condition` * `linear_algebra.matrix.charpoly.minpoly` * `linear_algebra.matrix.special_linear_group` * `linear_algebra.trace` * `measure_theory.function.strongly_measurable.basic` * `representation_theory.basic` * `ring_theory.fractional_ideal` * `ring_theory.hahn_series` * `ring_theory.henselian` * `ring_theory.ideal.over` * `ring_theory.localization.integral` * `ring_theory.power_basis` * `topology.instances.irrational` * `topology.sheaves.sheaf_condition.opens_le_cover` --- src/algebra/category/Ring/colimits.lean | 3 +++ src/analysis/box_integral/partition/additive.lean | 3 +++ src/analysis/calculus/formal_multilinear_series.lean | 3 +++ src/analysis/convex/cone/basic.lean | 3 +++ src/analysis/convex/side.lean | 3 +++ src/data/nat/prime_norm_num.lean | 3 +++ src/field_theory/ax_grothendieck.lean | 3 +++ src/field_theory/intermediate_field.lean | 3 +++ src/field_theory/separable.lean | 3 +++ src/linear_algebra/annihilating_polynomial.lean | 3 +++ src/linear_algebra/charpoly/basic.lean | 3 +++ src/linear_algebra/free_module/strong_rank_condition.lean | 3 +++ src/linear_algebra/matrix/charpoly/minpoly.lean | 3 +++ src/linear_algebra/matrix/special_linear_group.lean | 3 +++ src/linear_algebra/trace.lean | 3 +++ src/measure_theory/function/strongly_measurable/basic.lean | 3 +++ src/representation_theory/basic.lean | 3 +++ src/ring_theory/fractional_ideal.lean | 3 +++ src/ring_theory/hahn_series.lean | 3 +++ src/ring_theory/henselian.lean | 3 +++ src/ring_theory/ideal/over.lean | 3 +++ src/ring_theory/localization/integral.lean | 3 +++ src/ring_theory/power_basis.lean | 3 +++ src/topology/instances/irrational.lean | 3 +++ src/topology/sheaves/sheaf_condition/opens_le_cover.lean | 3 +++ 25 files changed, 75 insertions(+) diff --git a/src/algebra/category/Ring/colimits.lean b/src/algebra/category/Ring/colimits.lean index 1d3bf64ea8c69..c14f79f0734f1 100644 --- a/src/algebra/category/Ring/colimits.lean +++ b/src/algebra/category/Ring/colimits.lean @@ -10,6 +10,9 @@ import category_theory.concrete_category.elementwise /-! # The category of commutative rings has all colimits. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file uses a "pre-automated" approach, just as for `Mon/colimits.lean`. It is a very uniform approach, that conceivably could be synthesised directly by a tactic that analyses the shape of `comm_ring` and `ring_hom`. diff --git a/src/analysis/box_integral/partition/additive.lean b/src/analysis/box_integral/partition/additive.lean index 96594987e60c4..4932c78468f17 100644 --- a/src/analysis/box_integral/partition/additive.lean +++ b/src/analysis/box_integral/partition/additive.lean @@ -9,6 +9,9 @@ import analysis.normed_space.operator_norm /-! # Box additive functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We say that a function `f : box ι → M` from boxes in `ℝⁿ` to a commutative additive monoid `M` is *box additive* on subboxes of `I₀ : with_top (box ι)` if for any box `J`, `↑J ≤ I₀`, and a partition `π` of `J`, `f J = ∑ J' in π.boxes, f J'`. We use `I₀ : with_top (box ι)` instead of `I₀ : box ι` to diff --git a/src/analysis/calculus/formal_multilinear_series.lean b/src/analysis/calculus/formal_multilinear_series.lean index 36c44df149a6e..78aba40aa4ea6 100644 --- a/src/analysis/calculus/formal_multilinear_series.lean +++ b/src/analysis/calculus/formal_multilinear_series.lean @@ -8,6 +8,9 @@ import analysis.normed_space.multilinear /-! # Formal multilinear series +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define `formal_multilinear_series 𝕜 E F` to be a family of `n`-multilinear maps for all `n`, designed to model the sequence of derivatives of a function. In other files we use this notion to define `C^n` functions (called `cont_diff` in `mathlib`) and analytic functions. diff --git a/src/analysis/convex/cone/basic.lean b/src/analysis/convex/cone/basic.lean index ef9525cbd03ff..3feb656d347f0 100644 --- a/src/analysis/convex/cone/basic.lean +++ b/src/analysis/convex/cone/basic.lean @@ -10,6 +10,9 @@ import linear_algebra.linear_pmap /-! # Convex cones +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In a `𝕜`-module `E`, we define a convex cone as a set `s` such that `a • x + b • y ∈ s` whenever `x, y ∈ s` and `a, b > 0`. We prove that convex cones form a `complete_lattice`, and define their images (`convex_cone.map`) and preimages (`convex_cone.comap`) under linear maps. diff --git a/src/analysis/convex/side.lean b/src/analysis/convex/side.lean index a6941b8d4cc43..b4f12bc9539d7 100644 --- a/src/analysis/convex/side.lean +++ b/src/analysis/convex/side.lean @@ -10,6 +10,9 @@ import analysis.normed.group.add_torsor /-! # Sides of affine subspaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines notions of two points being on the same or opposite sides of an affine subspace. ## Main definitions diff --git a/src/data/nat/prime_norm_num.lean b/src/data/nat/prime_norm_num.lean index a5768ccb6f6ce..664fd3f285cb9 100644 --- a/src/data/nat/prime_norm_num.lean +++ b/src/data/nat/prime_norm_num.lean @@ -10,6 +10,9 @@ import tactic.norm_num /-! # Primality prover +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file provides a `norm_num` extention to prove that natural numbers are prime. -/ diff --git a/src/field_theory/ax_grothendieck.lean b/src/field_theory/ax_grothendieck.lean index 7cec7336d2e9f..691f6f9a276b4 100644 --- a/src/field_theory/ax_grothendieck.lean +++ b/src/field_theory/ax_grothendieck.lean @@ -10,6 +10,9 @@ import data.fintype.card /-! # Ax-Grothendieck for algebraic extensions of `zmod p` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves that if `R` is an algebraic extension of a finite field, then any injective polynomial map `R^n -> R^n` is also surjective. diff --git a/src/field_theory/intermediate_field.lean b/src/field_theory/intermediate_field.lean index de0b97b760c5a..241d7b54f2b42 100644 --- a/src/field_theory/intermediate_field.lean +++ b/src/field_theory/intermediate_field.lean @@ -11,6 +11,9 @@ import field_theory.tower /-! # Intermediate fields +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Let `L / K` be a field extension, given as an instance `algebra K L`. This file defines the type of fields in between `K` and `L`, `intermediate_field K L`. An `intermediate_field K L` is a subfield of `L` which contains (the image of) `K`, diff --git a/src/field_theory/separable.lean b/src/field_theory/separable.lean index af29f7b6c2003..6ba14bbcdd1ea 100644 --- a/src/field_theory/separable.lean +++ b/src/field_theory/separable.lean @@ -14,6 +14,9 @@ import ring_theory.power_basis # Separable polynomials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define a polynomial to be separable if it is coprime with its derivative. We prove basic properties about separable polynomials here. diff --git a/src/linear_algebra/annihilating_polynomial.lean b/src/linear_algebra/annihilating_polynomial.lean index 784fee1a4b31b..26a8531effd23 100644 --- a/src/linear_algebra/annihilating_polynomial.lean +++ b/src/linear_algebra/annihilating_polynomial.lean @@ -9,6 +9,9 @@ import ring_theory.principal_ideal_domain /-! # Annihilating Ideal +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given a commutative ring `R` and an `R`-algebra `A` Every element `a : A` defines an ideal `polynomial.ann_ideal a ⊆ R[X]`. diff --git a/src/linear_algebra/charpoly/basic.lean b/src/linear_algebra/charpoly/basic.lean index 8941aab44f79e..7fc3fd13a2b60 100644 --- a/src/linear_algebra/charpoly/basic.lean +++ b/src/linear_algebra/charpoly/basic.lean @@ -12,6 +12,9 @@ import field_theory.minpoly.field # Characteristic polynomial +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define the characteristic polynomial of `f : M →ₗ[R] M`, where `M` is a finite and free `R`-module. The proof that `f.charpoly` is the characteristic polynomial of the matrix of `f` in any basis is in `linear_algebra/charpoly/to_matrix`. diff --git a/src/linear_algebra/free_module/strong_rank_condition.lean b/src/linear_algebra/free_module/strong_rank_condition.lean index 9b5a50c17bdff..26fa94a29741c 100644 --- a/src/linear_algebra/free_module/strong_rank_condition.lean +++ b/src/linear_algebra/free_module/strong_rank_condition.lean @@ -11,6 +11,9 @@ import linear_algebra.invariant_basis_number # Strong rank condition for commutative rings +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We prove that any nontrivial commutative ring satisfies `strong_rank_condition`, meaning that if there is an injective linear map `(fin n → R) →ₗ[R] fin m → R`, then `n ≤ m`. This implies that any commutative ring satisfies `invariant_basis_number`: the rank of a finitely generated free diff --git a/src/linear_algebra/matrix/charpoly/minpoly.lean b/src/linear_algebra/matrix/charpoly/minpoly.lean index f8d20bd6c3433..ca0ffecda79b6 100644 --- a/src/linear_algebra/matrix/charpoly/minpoly.lean +++ b/src/linear_algebra/matrix/charpoly/minpoly.lean @@ -11,6 +11,9 @@ import ring_theory.power_basis /-! # The minimal polynomial divides the characteristic polynomial of a matrix. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This also includes some miscellaneous results about `minpoly` on matrices. -/ diff --git a/src/linear_algebra/matrix/special_linear_group.lean b/src/linear_algebra/matrix/special_linear_group.lean index 82548fa6bd476..1fcb0587bf398 100644 --- a/src/linear_algebra/matrix/special_linear_group.lean +++ b/src/linear_algebra/matrix/special_linear_group.lean @@ -10,6 +10,9 @@ import linear_algebra.matrix.to_lin /-! # The Special Linear group $SL(n, R)$ +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the elements of the Special Linear group `special_linear_group n R`, consisting of all square `R`-matrices with determinant `1` on the fintype `n` by `n`. In addition, we define the group structure on `special_linear_group n R` and the embedding into the general linear group diff --git a/src/linear_algebra/trace.lean b/src/linear_algebra/trace.lean index c164ee18444f2..c7e2673e6080a 100644 --- a/src/linear_algebra/trace.lean +++ b/src/linear_algebra/trace.lean @@ -14,6 +14,9 @@ import linear_algebra.projection /-! # Trace of a linear map +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the trace of a linear map. See also `linear_algebra/matrix/trace.lean` for the trace of a matrix. diff --git a/src/measure_theory/function/strongly_measurable/basic.lean b/src/measure_theory/function/strongly_measurable/basic.lean index df0f43cb089d0..6315583d2588b 100644 --- a/src/measure_theory/function/strongly_measurable/basic.lean +++ b/src/measure_theory/function/strongly_measurable/basic.lean @@ -12,6 +12,9 @@ import measure_theory.function.simple_func_dense /-! # Strongly measurable and finitely strongly measurable functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A function `f` is said to be strongly measurable if `f` is the sequential limit of simple functions. It is said to be finitely strongly measurable with respect to a measure `μ` if the supports of those simple functions have finite measure. We also provide almost everywhere versions of diff --git a/src/representation_theory/basic.lean b/src/representation_theory/basic.lean index 9bef318df98be..8d13447b4d399 100644 --- a/src/representation_theory/basic.lean +++ b/src/representation_theory/basic.lean @@ -13,6 +13,9 @@ import ring_theory.tensor_product /-! # Monoid representations +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file introduces monoid representations and their characters and defines a few ways to construct representations. diff --git a/src/ring_theory/fractional_ideal.lean b/src/ring_theory/fractional_ideal.lean index 676ec477949c0..7eea94685ba78 100644 --- a/src/ring_theory/fractional_ideal.lean +++ b/src/ring_theory/fractional_ideal.lean @@ -14,6 +14,9 @@ import tactic.field_simp /-! # Fractional ideals +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines fractional ideals of an integral domain and proves basic facts about them. ## Main definitions diff --git a/src/ring_theory/hahn_series.lean b/src/ring_theory/hahn_series.lean index 03abade5fa663..e0e25df39dd32 100644 --- a/src/ring_theory/hahn_series.lean +++ b/src/ring_theory/hahn_series.lean @@ -13,6 +13,9 @@ import algebra.order.group.with_top /-! # Hahn Series + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. If `Γ` is ordered and `R` has zero, then `hahn_series Γ R` consists of formal series over `Γ` with coefficients in `R`, whose supports are partially well-ordered. With further structure on `R` and `Γ`, we can add further structure on `hahn_series Γ R`, with the most studied case being when `Γ` is diff --git a/src/ring_theory/henselian.lean b/src/ring_theory/henselian.lean index 7e98fab4e198a..eb5ff41b66ed1 100644 --- a/src/ring_theory/henselian.lean +++ b/src/ring_theory/henselian.lean @@ -11,6 +11,9 @@ import linear_algebra.adic_completion /-! # Henselian rings +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we set up the basic theory of Henselian (local) rings. A ring `R` is *Henselian* at an ideal `I` if the following conditions hold: * `I` is contained in the Jacobson radical of `R` diff --git a/src/ring_theory/ideal/over.lean b/src/ring_theory/ideal/over.lean index ee5834597559e..e9a5a86776ecf 100644 --- a/src/ring_theory/ideal/over.lean +++ b/src/ring_theory/ideal/over.lean @@ -11,6 +11,9 @@ import ring_theory.localization.integral /-! # Ideals over/under ideals +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file concerns ideals lying over other ideals. Let `f : R →+* S` be a ring homomorphism (typically a ring extension), `I` an ideal of `R` and `J` an ideal of `S`. We say `J` lies over `I` (and `I` under `J`) if `I` is the `f`-preimage of `J`. diff --git a/src/ring_theory/localization/integral.lean b/src/ring_theory/localization/integral.lean index 0f9d2fd5a2851..3ffbf6a9d4928 100644 --- a/src/ring_theory/localization/integral.lean +++ b/src/ring_theory/localization/integral.lean @@ -16,6 +16,9 @@ import tactic.ring_exp /-! # Integral and algebraic elements of a fraction field +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Implementation notes See `src/ring_theory/localization/basic.lean` for a design overview. diff --git a/src/ring_theory/power_basis.lean b/src/ring_theory/power_basis.lean index bece1cd0a14ef..8cb2b73074fd1 100644 --- a/src/ring_theory/power_basis.lean +++ b/src/ring_theory/power_basis.lean @@ -8,6 +8,9 @@ import field_theory.minpoly.field /-! # Power basis +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines a structure `power_basis R S`, giving a basis of the `R`-algebra `S` as a finite list of powers `1, x, ..., x^n`. For example, if `x` is algebraic over a ring/field, adjoining `x` diff --git a/src/topology/instances/irrational.lean b/src/topology/instances/irrational.lean index 357cfcadc2810..dade2c4757ce0 100644 --- a/src/topology/instances/irrational.lean +++ b/src/topology/instances/irrational.lean @@ -10,6 +10,9 @@ import topology.metric_space.baire /-! # Topology of irrational numbers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove the following theorems: * `is_Gδ_irrational`, `dense_irrational`, `eventually_residual_irrational`: irrational numbers diff --git a/src/topology/sheaves/sheaf_condition/opens_le_cover.lean b/src/topology/sheaves/sheaf_condition/opens_le_cover.lean index 59e5d0758b65f..c74bb17306eea 100644 --- a/src/topology/sheaves/sheaf_condition/opens_le_cover.lean +++ b/src/topology/sheaves/sheaf_condition/opens_le_cover.lean @@ -8,6 +8,9 @@ import topology.sheaves.sheaf_condition.sites /-! # Another version of the sheaf condition. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given a family of open sets `U : ι → opens X` we can form the subcategory `{ V : opens X // ∃ i, V ≤ U i }`, which has `supr U` as a cocone. From 6315581f5650ffa2fbdbbbedc41243c8d7070981 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Wed, 24 May 2023 11:30:13 +0000 Subject: [PATCH 1068/1291] feat(probability/kernel/disintegration): integral against `cond_kernel` (#19066) Co-authored-by: RemyDegenne --- src/probability/kernel/disintegration.lean | 147 +++++++++++++++++++++ 1 file changed, 147 insertions(+) diff --git a/src/probability/kernel/disintegration.lean b/src/probability/kernel/disintegration.lean index d361b61fc6414..a4cf308924105 100644 --- a/src/probability/kernel/disintegration.lean +++ b/src/probability/kernel/disintegration.lean @@ -5,6 +5,7 @@ Authors: Rémy Degenne -/ import probability.kernel.cond_cdf import measure_theory.constructions.polish +import probability.kernel.integral_comp_prod /-! # Disintegration of measures on product spaces @@ -398,6 +399,25 @@ begin simp_rw [kernel.comp_prod_apply _ _ _ hs, kernel.const_apply, kernel.prod_mk_left_apply], end +lemma set_lintegral_cond_kernel_eq_measure_prod {s : set α} (hs : measurable_set s) + {t : set Ω} (ht : measurable_set t) : + ∫⁻ a in s, ρ.cond_kernel a t ∂ρ.fst = ρ (s ×ˢ t) := +begin + have : ρ (s ×ˢ t) = ((kernel.const unit ρ.fst ⊗ₖ kernel.prod_mk_left unit ρ.cond_kernel) ()) + (s ×ˢ t), + { congr, exact measure_eq_comp_prod ρ, }, + rw [this, kernel.comp_prod_apply _ _ _ (hs.prod ht)], + simp only [prod_mk_mem_set_prod_eq, kernel.lintegral_const, kernel.prod_mk_left_apply], + rw ← lintegral_indicator _ hs, + congr, + ext1 x, + classical, + rw indicator_apply, + split_ifs with hx, + { simp only [hx, if_true, true_and, set_of_mem_eq], }, + { simp only [hx, if_false, false_and, set_of_false, measure_empty], }, +end + lemma lintegral_cond_kernel {f : α × Ω → ℝ≥0∞} (hf : measurable f) : ∫⁻ a, ∫⁻ ω, f (a, ω) ∂(ρ.cond_kernel a) ∂ρ.fst = ∫⁻ x, f x ∂ρ := begin @@ -406,6 +426,133 @@ begin simp_rw kernel.prod_mk_left_apply, end +lemma set_lintegral_cond_kernel {f : α × Ω → ℝ≥0∞} (hf : measurable f) + {s : set α} (hs : measurable_set s) {t : set Ω} (ht : measurable_set t) : + ∫⁻ a in s, ∫⁻ ω in t, f (a, ω) ∂(ρ.cond_kernel a) ∂ρ.fst = ∫⁻ x in s ×ˢ t, f x ∂ρ := +begin + conv_rhs { rw measure_eq_comp_prod ρ, }, + simp_rw [← kernel.restrict_apply _ (hs.prod ht), ← kernel.comp_prod_restrict, + kernel.lintegral_comp_prod _ _ _ hf, kernel.restrict_apply, kernel.const_apply, + kernel.prod_mk_left_apply], +end + +lemma set_lintegral_cond_kernel_univ_right {f : α × Ω → ℝ≥0∞} (hf : measurable f) + {s : set α} (hs : measurable_set s) : + ∫⁻ a in s, ∫⁻ ω, f (a, ω) ∂(ρ.cond_kernel a) ∂ρ.fst = ∫⁻ x in s ×ˢ univ, f x ∂ρ := +by { rw ← set_lintegral_cond_kernel ρ hf hs measurable_set.univ, simp_rw measure.restrict_univ, } + +lemma set_lintegral_cond_kernel_univ_left {f : α × Ω → ℝ≥0∞} (hf : measurable f) + {t : set Ω} (ht : measurable_set t) : + ∫⁻ a, ∫⁻ ω in t, f (a, ω) ∂(ρ.cond_kernel a) ∂ρ.fst = ∫⁻ x in univ ×ˢ t, f x ∂ρ := +by { rw ← set_lintegral_cond_kernel ρ hf measurable_set.univ ht, simp_rw measure.restrict_univ, } + +section integral_cond_kernel + +variables {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] [complete_space E] + +lemma _root_.measure_theory.ae_strongly_measurable.integral_cond_kernel + {ρ : measure (α × Ω)} [is_finite_measure ρ] {f : α × Ω → E} (hf : ae_strongly_measurable f ρ) : + ae_strongly_measurable (λ x, ∫ y, f (x, y) ∂(ρ.cond_kernel x)) ρ.fst := +begin + rw measure_eq_comp_prod ρ at hf, + exact ae_strongly_measurable.integral_kernel_comp_prod hf, +end + +lemma integral_cond_kernel {ρ : measure (α × Ω)} [is_finite_measure ρ] + {f : α × Ω → E} (hf : integrable f ρ) : + ∫ a, ∫ x, f (a, x) ∂(ρ.cond_kernel a) ∂ρ.fst = ∫ ω, f ω ∂ρ := +begin + conv_rhs { rw measure_eq_comp_prod ρ, }, + have hf' : integrable f ((kernel.const unit ρ.fst ⊗ₖ kernel.prod_mk_left unit ρ.cond_kernel) ()), + { rwa measure_eq_comp_prod ρ at hf, }, + rw [integral_comp_prod hf', kernel.const_apply], + simp_rw kernel.prod_mk_left_apply, +end + +lemma set_integral_cond_kernel {ρ : measure (α × Ω)} [is_finite_measure ρ] + {f : α × Ω → E} {s : set α} (hs : measurable_set s) + {t : set Ω} (ht : measurable_set t) (hf : integrable_on f (s ×ˢ t) ρ) : + ∫ a in s, ∫ ω in t, f (a, ω) ∂(ρ.cond_kernel a) ∂ρ.fst = ∫ x in s ×ˢ t, f x ∂ρ := +begin + conv_rhs { rw measure_eq_comp_prod ρ, }, + rw set_integral_comp_prod hs ht, + { simp_rw [kernel.prod_mk_left_apply, kernel.const_apply], }, + { rwa measure_eq_comp_prod ρ at hf, }, +end + +lemma set_integral_cond_kernel_univ_right {ρ : measure (α × Ω)} [is_finite_measure ρ] + {f : α × Ω → E} {s : set α} (hs : measurable_set s) + (hf : integrable_on f (s ×ˢ univ) ρ) : + ∫ a in s, ∫ ω, f (a, ω) ∂(ρ.cond_kernel a) ∂ρ.fst = ∫ x in s ×ˢ univ, f x ∂ρ := +by { rw ← set_integral_cond_kernel hs measurable_set.univ hf, simp_rw measure.restrict_univ, } + +lemma set_integral_cond_kernel_univ_left {ρ : measure (α × Ω)} [is_finite_measure ρ] + {f : α × Ω → E} {t : set Ω} (ht : measurable_set t) + (hf : integrable_on f (univ ×ˢ t) ρ) : + ∫ a, ∫ ω in t, f (a, ω) ∂(ρ.cond_kernel a) ∂ρ.fst = ∫ x in univ ×ˢ t, f x ∂ρ := +by { rw ← set_integral_cond_kernel measurable_set.univ ht hf, simp_rw measure.restrict_univ, } + +end integral_cond_kernel + end polish end probability_theory + +namespace measure_theory +/-! ### Integrability + +We place these lemmas in the `measure_theory` namespace to enable dot notation. -/ + +open probability_theory + +variables {α Ω E F : Type*} {mα : measurable_space α} [measurable_space Ω] [topological_space Ω] + [borel_space Ω] [polish_space Ω] [nonempty Ω] + [normed_add_comm_group E] [normed_space ℝ E] [complete_space E] + [normed_add_comm_group F] + {ρ : measure (α × Ω)} [is_finite_measure ρ] + +include mα + +lemma ae_strongly_measurable.ae_integrable_cond_kernel_iff + {f : α × Ω → F} (hf : ae_strongly_measurable f ρ) : + (∀ᵐ a ∂ρ.fst, integrable (λ ω, f (a, ω)) (ρ.cond_kernel a)) + ∧ integrable (λ a, ∫ ω, ‖f (a, ω)‖ ∂(ρ.cond_kernel a)) ρ.fst + ↔ integrable f ρ := +begin + rw measure_eq_comp_prod ρ at hf, + conv_rhs { rw measure_eq_comp_prod ρ, }, + rw integrable_comp_prod_iff hf, + simp_rw [kernel.prod_mk_left_apply, kernel.const_apply], +end + +lemma integrable.cond_kernel_ae {f : α × Ω → F} (hf_int : integrable f ρ) : + ∀ᵐ a ∂ρ.fst, integrable (λ ω, f (a, ω)) (ρ.cond_kernel a) := +begin + have hf_ae : ae_strongly_measurable f ρ := hf_int.1, + rw ← hf_ae.ae_integrable_cond_kernel_iff at hf_int, + exact hf_int.1, +end + +lemma integrable.integral_norm_cond_kernel {f : α × Ω → F} (hf_int : integrable f ρ) : + integrable (λ x, ∫ y, ‖f (x, y)‖ ∂(ρ.cond_kernel x)) ρ.fst := +begin + have hf_ae : ae_strongly_measurable f ρ := hf_int.1, + rw ← hf_ae.ae_integrable_cond_kernel_iff at hf_int, + exact hf_int.2, +end + +lemma integrable.norm_integral_cond_kernel {f : α × Ω → E} (hf_int : integrable f ρ) : + integrable (λ x, ‖∫ y, f (x, y) ∂(ρ.cond_kernel x)‖) ρ.fst := +begin + refine hf_int.integral_norm_cond_kernel.mono hf_int.1.integral_cond_kernel.norm _, + refine eventually_of_forall (λ x, _), + rw norm_norm, + refine (norm_integral_le_integral_norm _).trans_eq (real.norm_of_nonneg _).symm, + exact integral_nonneg_of_ae (eventually_of_forall (λ y, norm_nonneg _)), +end + +lemma integrable.integral_cond_kernel {f : α × Ω → E} (hf_int : integrable f ρ) : + integrable (λ x, ∫ y, f (x, y) ∂(ρ.cond_kernel x)) ρ.fst := +(integrable_norm_iff hf_int.1.integral_cond_kernel).mp hf_int.norm_integral_cond_kernel + +end measure_theory From bc91ed7093bf098d253401e69df601fc33dde156 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 24 May 2023 12:44:51 +0000 Subject: [PATCH 1069/1291] feat(analysis/normed_space/basic): scaling a set scales its diameter, translating it leaves it unchanged (#18990) --- src/analysis/normed/mul_action.lean | 11 ++-- src/analysis/normed_space/basic.lean | 2 +- src/analysis/normed_space/pointwise.lean | 54 ++++++++++++++++++- src/data/real/ennreal.lean | 10 +++- .../metric_space/hausdorff_distance.lean | 8 ++- src/topology/metric_space/isometric_smul.lean | 9 ++++ 6 files changed, 87 insertions(+), 7 deletions(-) diff --git a/src/analysis/normed/mul_action.lean b/src/analysis/normed/mul_action.lean index 91e0cc442be98..5106c76fea567 100644 --- a/src/analysis/normed/mul_action.lean +++ b/src/analysis/normed/mul_action.lean @@ -34,7 +34,10 @@ by simpa only [dist_eq_norm, sub_zero] using dist_smul_pair s x y lemma nndist_smul_le (s : α) (x y : β) : nndist (s • x) (s • y) ≤ ‖s‖₊ * nndist x y := dist_smul_le s x y -lemma lipschitz_with_smul (s : α) : lipschitz_with ‖s‖₊ ((•) s : β → β) := +lemma edist_smul_le (s : α) (x y : β) : edist (s • x) (s • y) ≤ ‖s‖₊ • edist x y := +by simpa only [edist_nndist, ennreal.coe_mul] using ennreal.coe_le_coe.mpr (nndist_smul_le s x y) + +lemma lipschitz_with_smul (s : α) : lipschitz_with ‖s‖₊ ((•) s : β → β) := lipschitz_with_iff_dist_le_mul.2 $ dist_smul_le _ end seminormed_add_group @@ -91,8 +94,10 @@ variables [module α β] [has_bounded_smul α β] lemma dist_smul₀ (s : α) (x y : β) : dist (s • x) (s • y) = ‖s‖ * dist x y := by simp_rw [dist_eq_norm, (norm_smul _ _).symm, smul_sub] -lemma nndist_smul₀ (s : α) (x y : β) : - nndist (s • x) (s • y) = ‖s‖₊ * nndist x y := +lemma nndist_smul₀ (s : α) (x y : β) : nndist (s • x) (s • y) = ‖s‖₊ * nndist x y := nnreal.eq $ dist_smul₀ s x y +lemma edist_smul₀ (s : α) (x y : β) : edist (s • x) (s • y) = ‖s‖₊ • edist x y := +by simp only [edist_nndist, nndist_smul₀, ennreal.coe_mul, ennreal.smul_def, smul_eq_mul] + end normed_division_ring_module diff --git a/src/analysis/normed_space/basic.lean b/src/analysis/normed_space/basic.lean index b75f459755dcf..6b5da2ca732c5 100644 --- a/src/analysis/normed_space/basic.lean +++ b/src/analysis/normed_space/basic.lean @@ -23,7 +23,7 @@ about these definitions. variables {α : Type*} {β : Type*} {γ : Type*} {ι : Type*} open filter metric function set -open_locale topology big_operators nnreal ennreal uniformity pointwise +open_locale topology big_operators nnreal ennreal uniformity section seminormed_add_comm_group diff --git a/src/analysis/normed_space/pointwise.lean b/src/analysis/normed_space/pointwise.lean index 65e8d118b8b47..adb409acbb151 100644 --- a/src/analysis/normed_space/pointwise.lean +++ b/src/analysis/normed_space/pointwise.lean @@ -21,7 +21,59 @@ multiplication of bounded sets remain bounded. open metric set open_locale pointwise topology -variables {𝕜 E : Type*} [normed_field 𝕜] +variables {𝕜 E : Type*} + +section smul_zero_class +variables [seminormed_add_comm_group 𝕜] [seminormed_add_comm_group E] +variables [smul_zero_class 𝕜 E] [has_bounded_smul 𝕜 E] + +lemma ediam_smul_le (c : 𝕜) (s : set E) : + emetric.diam (c • s) ≤ ‖c‖₊ • emetric.diam s := +(lipschitz_with_smul c).ediam_image_le s + +end smul_zero_class + +section division_ring +variables [normed_division_ring 𝕜] [seminormed_add_comm_group E] +variables [module 𝕜 E] [has_bounded_smul 𝕜 E] + +lemma ediam_smul₀ (c : 𝕜) (s : set E) : + emetric.diam (c • s) = ‖c‖₊ • emetric.diam s := +begin + refine le_antisymm (ediam_smul_le c s) _, + obtain rfl | hc := eq_or_ne c 0, + { obtain rfl | hs := s.eq_empty_or_nonempty, + { simp }, + simp [zero_smul_set hs, ←set.singleton_zero], }, + { have := (lipschitz_with_smul c⁻¹).ediam_image_le (c • s), + rwa [← smul_eq_mul, ←ennreal.smul_def, set.image_smul, inv_smul_smul₀ hc s, nnnorm_inv, + ennreal.le_inv_smul_iff (nnnorm_ne_zero_iff.mpr hc)] at this } +end + +lemma diam_smul₀ (c : 𝕜) (x : set E) : diam (c • x) = ‖c‖ * diam x := +by simp_rw [diam, ediam_smul₀, ennreal.to_real_smul, nnreal.smul_def, coe_nnnorm, smul_eq_mul] + +lemma inf_edist_smul₀ {c : 𝕜} (hc : c ≠ 0) (s : set E) (x : E) : + emetric.inf_edist (c • x) (c • s) = ‖c‖₊ • emetric.inf_edist x s := +begin + simp_rw [emetric.inf_edist], + have : function.surjective ((•) c : E → E) := + function.right_inverse.surjective (smul_inv_smul₀ hc), + transitivity ⨅ y (H : y ∈ s), ‖c‖₊ • edist x y, + { refine (this.infi_congr _ $ λ y, _).symm, + simp_rw [smul_mem_smul_set_iff₀ hc, edist_smul₀] }, + { have : (‖c‖₊ : ennreal) ≠ 0 := by simp [hc], + simp_rw [ennreal.smul_def, smul_eq_mul, ennreal.mul_infi_of_ne this ennreal.coe_ne_top] }, +end + +lemma inf_dist_smul₀ {c : 𝕜} (hc : c ≠ 0) (s : set E) (x : E) : + metric.inf_dist (c • x) (c • s) = ‖c‖ * metric.inf_dist x s := +by simp_rw [metric.inf_dist, inf_edist_smul₀ hc, ennreal.to_real_smul, nnreal.smul_def, coe_nnnorm, + smul_eq_mul] + +end division_ring + +variables [normed_field 𝕜] section seminormed_add_comm_group variables [seminormed_add_comm_group E] [normed_space 𝕜 E] diff --git a/src/data/real/ennreal.lean b/src/data/real/ennreal.lean index fda281f3e2bfe..f54666b723bf6 100644 --- a/src/data/real/ennreal.lean +++ b/src/data/real/ennreal.lean @@ -1243,9 +1243,17 @@ begin simpa [left_ne_zero_of_mul_eq_one h] using h, end -lemma mul_le_iff_le_inv {a b r : ℝ≥0∞} (hr₀ : r ≠ 0) (hr₁ : r ≠ ∞) : (r * a ≤ b ↔ a ≤ r⁻¹ * b) := +lemma mul_le_iff_le_inv {a b r : ℝ≥0∞} (hr₀ : r ≠ 0) (hr₁ : r ≠ ∞) : r * a ≤ b ↔ a ≤ r⁻¹ * b := by rw [←@ennreal.mul_le_mul_left _ a _ hr₀ hr₁, ←mul_assoc, ennreal.mul_inv_cancel hr₀ hr₁, one_mul] +/-- A variant of `le_inv_smul_iff` that holds for `ennreal`. -/ +protected lemma le_inv_smul_iff {a b : ℝ≥0∞} {r : ℝ≥0} (hr₀ : r ≠ 0) : a ≤ r⁻¹ • b ↔ r • a ≤ b := +by simpa [hr₀, ennreal.smul_def] using (mul_le_iff_le_inv (coe_ne_zero.mpr hr₀) coe_ne_top).symm + +/-- A variant of `inv_smul_le_iff` that holds for `ennreal`. -/ +protected lemma inv_smul_le_iff {a b : ℝ≥0∞} {r : ℝ≥0} (hr₀ : r ≠ 0) : r⁻¹ • a ≤ b ↔ a ≤ r • b := +by simpa only [inv_inv] using (ennreal.le_inv_smul_iff (inv_ne_zero hr₀)).symm + lemma le_of_forall_nnreal_lt {x y : ℝ≥0∞} (h : ∀ r : ℝ≥0, ↑r < x → ↑r ≤ y) : x ≤ y := begin refine le_of_forall_ge_of_dense (λ r hr, _), diff --git a/src/topology/metric_space/hausdorff_distance.lean b/src/topology/metric_space/hausdorff_distance.lean index 39a3c9bc4a3ee..a8150266b51aa 100644 --- a/src/topology/metric_space/hausdorff_distance.lean +++ b/src/topology/metric_space/hausdorff_distance.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ import analysis.specific_limits.basic +import topology.metric_space.isometric_smul import topology.metric_space.isometry import topology.instances.ennreal @@ -30,7 +31,7 @@ This files introduces: * `cthickening δ s`, the closed thickening by radius `δ` of a set `s` in a pseudo emetric space. -/ noncomputable theory -open_locale classical nnreal ennreal topology +open_locale classical nnreal ennreal topology pointwise universes u v w open classical set function topological_space filter @@ -167,6 +168,11 @@ lemma inf_edist_image (hΦ : isometry Φ) : inf_edist (Φ x) (Φ '' t) = inf_edist x t := by simp only [inf_edist, infi_image, hΦ.edist_eq] +@[simp, to_additive] lemma inf_edist_smul {M} [has_smul M α] [has_isometric_smul M α] + (c : M) (x : α) (s : set α) : + inf_edist (c • x) (c • s) = inf_edist x s := +inf_edist_image (isometry_smul _ _) + lemma _root_.is_open.exists_Union_is_closed {U : set α} (hU : is_open U) : ∃ F : ℕ → set α, (∀ n, is_closed (F n)) ∧ (∀ n, F n ⊆ U) ∧ ((⋃ n, F n) = U) ∧ monotone F := begin diff --git a/src/topology/metric_space/isometric_smul.lean b/src/topology/metric_space/isometric_smul.lean index 01e41594ac879..7259cf14344de 100644 --- a/src/topology/metric_space/isometric_smul.lean +++ b/src/topology/metric_space/isometric_smul.lean @@ -67,6 +67,10 @@ variables [pseudo_emetric_space X] [group G] [mul_action G X] [has_isometric_smu edist (c • x) (c • y) = edist x y := isometry_smul X c x y +@[simp, to_additive] lemma ediam_smul [has_smul M X] [has_isometric_smul M X] (c : M) (s : set X) : + emetric.diam (c • s) = emetric.diam s := +(isometry_smul _ _).ediam_image s + @[to_additive] lemma isometry_mul_left [has_mul M] [pseudo_emetric_space M] [has_isometric_smul M M] (a : M) : isometry ((*) a) := isometry_smul M a @@ -223,6 +227,11 @@ lemma nndist_smul [pseudo_metric_space X] [has_smul M X] [has_isometric_smul M X (c : M) (x y : X) : nndist (c • x) (c • y) = nndist x y := (isometry_smul X c).nndist_eq x y +@[simp, to_additive] +lemma diam_smul [pseudo_metric_space X] [has_smul M X] [has_isometric_smul M X] + (c : M) (s : set X) : metric.diam (c • s) = metric.diam s := +(isometry_smul _ _).diam_image s + @[simp, to_additive] lemma dist_mul_left [pseudo_metric_space M] [has_mul M] [has_isometric_smul M M] (a b c : M) : dist (a * b) (a * c) = dist b c := From 5bfbcca0a7ffdd21cf1682e59106d6c942434a32 Mon Sep 17 00:00:00 2001 From: Riccardo Brasca Date: Wed, 24 May 2023 14:35:21 +0000 Subject: [PATCH 1070/1291] chore(ring_theory.polynomial.cyclotomic): split file (#19077) `ring_theory.polynomial.cyclotomic` is almost 1000 lines long and it can be nicely split. We also fix some docstring. --- .../37_solution_of_cubic.lean | 2 +- src/number_theory/cyclotomic/basic.lean | 8 +- .../cyclotomic/primitive_roots.lean | 1 + .../polynomial/cyclotomic/basic.lean | 370 +----------------- .../polynomial/cyclotomic/eval.lean | 2 +- .../polynomial/cyclotomic/expand.lean | 186 +++++++++ .../polynomial/cyclotomic/roots.lean | 243 ++++++++++++ .../polynomial/eisenstein/is_integral.lean | 2 +- 8 files changed, 442 insertions(+), 372 deletions(-) create mode 100644 src/ring_theory/polynomial/cyclotomic/expand.lean create mode 100644 src/ring_theory/polynomial/cyclotomic/roots.lean diff --git a/archive/100-theorems-list/37_solution_of_cubic.lean b/archive/100-theorems-list/37_solution_of_cubic.lean index efa87d1353b09..aca8b61c940bc 100644 --- a/archive/100-theorems-list/37_solution_of_cubic.lean +++ b/archive/100-theorems-list/37_solution_of_cubic.lean @@ -5,7 +5,7 @@ Authors: Jeoff Lee -/ import tactic.linear_combination import ring_theory.roots_of_unity -import ring_theory.polynomial.cyclotomic.basic +import ring_theory.polynomial.cyclotomic.roots /-! # The Solution of a Cubic diff --git a/src/number_theory/cyclotomic/basic.lean b/src/number_theory/cyclotomic/basic.lean index bc009f7e7731f..5a324524c74a3 100644 --- a/src/number_theory/cyclotomic/basic.lean +++ b/src/number_theory/cyclotomic/basic.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Riccardo Brasca. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Riccardo Brasca -/ -import ring_theory.polynomial.cyclotomic.basic +import ring_theory.polynomial.cyclotomic.roots import number_theory.number_field.basic import field_theory.galois @@ -374,7 +374,8 @@ begin map_cyclotomic, mem_roots (cyclotomic_ne_zero n B)] at hx, simp only [mem_singleton_iff, exists_eq_left, mem_set_of_eq], rw is_root_of_unity_iff n.pos, - exact ⟨n, nat.mem_divisors_self n n.ne_zero, hx⟩ }, + exact ⟨n, nat.mem_divisors_self n n.ne_zero, hx⟩, + all_goals { apply_instance } }, { simp only [mem_singleton_iff, exists_eq_left, mem_set_of_eq] at hx, obtain ⟨i, hin, rfl⟩ := hζ.eq_pow_of_pow_eq_one hx n.pos, refine set_like.mem_coe.2 (subalgebra.pow_mem _ (subset_adjoin _) _), @@ -393,7 +394,8 @@ begin rw is_root_of_unity_iff n.pos, refine ⟨n, nat.mem_divisors_self n n.ne_zero, _⟩, rwa [finset.mem_coe, multiset.mem_to_finset, - map_cyclotomic, mem_roots $ cyclotomic_ne_zero n B] at hx }, + map_cyclotomic, mem_roots $ cyclotomic_ne_zero n B] at hx, + all_goals { apply_instance } }, { simp only [mem_singleton_iff, exists_eq_left, mem_set_of_eq] at hx, simpa only [hx, multiset.mem_to_finset, finset.mem_coe, map_cyclotomic, mem_roots (cyclotomic_ne_zero n B)] using hζ.is_root_cyclotomic n.pos } diff --git a/src/number_theory/cyclotomic/primitive_roots.lean b/src/number_theory/cyclotomic/primitive_roots.lean index 11c48e29db1d8..971334e1bba6b 100644 --- a/src/number_theory/cyclotomic/primitive_roots.lean +++ b/src/number_theory/cyclotomic/primitive_roots.lean @@ -10,6 +10,7 @@ import number_theory.cyclotomic.basic import ring_theory.adjoin.power_basis import ring_theory.polynomial.cyclotomic.eval import ring_theory.norm +import ring_theory.polynomial.cyclotomic.expand /-! # Primitive roots in cyclotomic fields diff --git a/src/ring_theory/polynomial/cyclotomic/basic.lean b/src/ring_theory/polynomial/cyclotomic/basic.lean index 90332b2db71e3..da333193aa6e1 100644 --- a/src/ring_theory/polynomial/cyclotomic/basic.lean +++ b/src/ring_theory/polynomial/cyclotomic/basic.lean @@ -30,13 +30,11 @@ with coefficients in any ring `R`. ## Main results -* `int_coeff_of_cycl` : If there is a primitive `n`-th root of unity in `K`, then `cyclotomic' n K` -comes from a polynomial with integer coefficients. -* `deg_of_cyclotomic` : The degree of `cyclotomic n` is `totient n`. -* `prod_cyclotomic_eq_X_pow_sub_one` : `X ^ n - 1 = ∏ (cyclotomic i)`, where `i` divides `n`. -* `cyclotomic_eq_prod_X_pow_sub_one_pow_moebius` : The Möbius inversion formula for +* `polynomial.degree_cyclotomic` : The degree of `cyclotomic n` is `totient n`. +* `polynomial.prod_cyclotomic_eq_X_pow_sub_one` : `X ^ n - 1 = ∏ (cyclotomic i)`, where `i` + divides `n`. +* `polynomial.cyclotomic_eq_prod_X_pow_sub_one_pow_moebius` : The Möbius inversion formula for `cyclotomic n R` over an abstract fraction field for `R[X]`. -* `cyclotomic.irreducible` : `cyclotomic n ℤ` is irreducible. ## Implementation details @@ -449,24 +447,6 @@ begin { exact λ hk, hdn.not_le $ nat.divisor_le hk } end -lemma _root_.is_root_of_unity_iff {n : ℕ} (h : 0 < n) (R : Type*) [comm_ring R] [is_domain R] - {ζ : R} : ζ ^ n = 1 ↔ ∃ i ∈ n.divisors, (cyclotomic i R).is_root ζ := -by rw [←mem_nth_roots h, nth_roots, mem_roots $ X_pow_sub_C_ne_zero h _, - C_1, ←prod_cyclotomic_eq_X_pow_sub_one h, is_root_prod]; apply_instance - -lemma is_root_of_unity_of_root_cyclotomic {n : ℕ} {R} [comm_ring R] {ζ : R} {i : ℕ} - (hi : i ∈ n.divisors) (h : (cyclotomic i R).is_root ζ) : ζ ^ n = 1 := -begin - rcases n.eq_zero_or_pos with rfl | hn, - { exact pow_zero _ }, - have := congr_arg (eval ζ) (prod_cyclotomic_eq_X_pow_sub_one hn R).symm, - rw [eval_sub, eval_pow, eval_X, eval_one] at this, - convert eq_add_of_sub_eq' this, - convert (add_zero _).symm, - apply eval_eq_zero_of_dvd_of_eval_eq_zero _ h, - exact finset.dvd_prod_of_mem _ hi -end - section arithmetic_function open nat.arithmetic_function open_locale arithmetic_function @@ -547,117 +527,6 @@ begin cyclotomic'_eq_X_pow_sub_one_div hpos hz, finset.prod_congr (refl k.proper_divisors) h] end -section roots - -variables {R : Type*} {n : ℕ} [comm_ring R] [is_domain R] - -/-- Any `n`-th primitive root of unity is a root of `cyclotomic n K`.-/ -lemma _root_.is_primitive_root.is_root_cyclotomic (hpos : 0 < n) {μ : R} - (h : is_primitive_root μ n) : is_root (cyclotomic n R) μ := -begin - rw [← mem_roots (cyclotomic_ne_zero n R), - cyclotomic_eq_prod_X_sub_primitive_roots h, roots_prod_X_sub_C, ← finset.mem_def], - rwa [← mem_primitive_roots hpos] at h, -end - -private lemma is_root_cyclotomic_iff' {n : ℕ} {K : Type*} [field K] {μ : K} [ne_zero (n : K)] : - is_root (cyclotomic n K) μ ↔ is_primitive_root μ n := -begin - -- in this proof, `o` stands for `order_of μ` - have hnpos : 0 < n := (ne_zero.of_ne_zero_coe K).out.bot_lt, - refine ⟨λ hμ, _, is_primitive_root.is_root_cyclotomic hnpos⟩, - have hμn : μ ^ n = 1, - { rw is_root_of_unity_iff hnpos, - exact ⟨n, n.mem_divisors_self hnpos.ne', hμ⟩ }, - by_contra hnμ, - have ho : 0 < order_of μ, - { apply order_of_pos', - rw is_of_fin_order_iff_pow_eq_one, - exact ⟨n, hnpos, hμn⟩ }, - have := pow_order_of_eq_one μ, - rw is_root_of_unity_iff ho at this, - obtain ⟨i, hio, hiμ⟩ := this, - replace hio := nat.dvd_of_mem_divisors hio, - rw is_primitive_root.not_iff at hnμ, - rw ←order_of_dvd_iff_pow_eq_one at hμn, - have key : i < n := (nat.le_of_dvd ho hio).trans_lt ((nat.le_of_dvd hnpos hμn).lt_of_ne hnμ), - have key' : i ∣ n := hio.trans hμn, - rw ←polynomial.dvd_iff_is_root at hμ hiμ, - have hni : {i, n} ⊆ n.divisors, - { simpa [finset.insert_subset, key'] using hnpos.ne' }, - obtain ⟨k, hk⟩ := hiμ, - obtain ⟨j, hj⟩ := hμ, - have := prod_cyclotomic_eq_X_pow_sub_one hnpos K, - rw [←finset.prod_sdiff hni, finset.prod_pair key.ne, hk, hj] at this, - have hn := (X_pow_sub_one_separable_iff.mpr $ ne_zero.nat_cast_ne n K).squarefree, - rw [←this, squarefree] at hn, - contrapose! hn, - refine ⟨X - C μ, ⟨(∏ x in n.divisors \ {i, n}, cyclotomic x K) * k * j, by ring⟩, _⟩, - simp [polynomial.is_unit_iff_degree_eq_zero] -end - -lemma is_root_cyclotomic_iff [ne_zero (n : R)] {μ : R} : - is_root (cyclotomic n R) μ ↔ is_primitive_root μ n := -begin - have hf : function.injective _ := is_fraction_ring.injective R (fraction_ring R), - haveI : ne_zero (n : fraction_ring R) := ne_zero.nat_of_injective hf, - rw [←is_root_map_iff hf, ←is_primitive_root.map_iff_of_injective hf, map_cyclotomic, - ←is_root_cyclotomic_iff'] -end - -lemma roots_cyclotomic_nodup [ne_zero (n : R)] : (cyclotomic n R).roots.nodup := -begin - obtain h | ⟨ζ, hζ⟩ := (cyclotomic n R).roots.empty_or_exists_mem, - { exact h.symm ▸ multiset.nodup_zero }, - rw [mem_roots $ cyclotomic_ne_zero n R, is_root_cyclotomic_iff] at hζ, - refine multiset.nodup_of_le (roots.le_of_dvd (X_pow_sub_C_ne_zero - (ne_zero.pos_of_ne_zero_coe R) 1) $ cyclotomic.dvd_X_pow_sub_one n R) hζ.nth_roots_nodup, -end - -lemma cyclotomic.roots_to_finset_eq_primitive_roots [ne_zero (n : R)] : - (⟨(cyclotomic n R).roots, roots_cyclotomic_nodup⟩ : finset _) = primitive_roots n R := -by { ext, simp [cyclotomic_ne_zero n R, is_root_cyclotomic_iff, - mem_primitive_roots, ne_zero.pos_of_ne_zero_coe R] } - -lemma cyclotomic.roots_eq_primitive_roots_val [ne_zero (n : R)] : - (cyclotomic n R).roots = (primitive_roots n R).val := -by rw ←cyclotomic.roots_to_finset_eq_primitive_roots - -end roots - -/-- If `R` is of characteristic zero, then `ζ` is a root of `cyclotomic n R` if and only if it is a -primitive `n`-th root of unity. -/ -lemma is_root_cyclotomic_iff_char_zero {n : ℕ} {R : Type*} [comm_ring R] [is_domain R] - [char_zero R] {μ : R} (hn : 0 < n) : - (polynomial.cyclotomic n R).is_root μ ↔ is_primitive_root μ n := -by { letI := ne_zero.of_gt hn, exact is_root_cyclotomic_iff } - -/-- Over a ring `R` of characteristic zero, `λ n, cyclotomic n R` is injective. -/ -lemma cyclotomic_injective {R : Type*} [comm_ring R] [char_zero R] : - function.injective (λ n, cyclotomic n R) := -begin - intros n m hnm, - simp only at hnm, - rcases eq_or_ne n 0 with rfl | hzero, - { rw [cyclotomic_zero] at hnm, - replace hnm := congr_arg nat_degree hnm, - rw [nat_degree_one, nat_degree_cyclotomic] at hnm, - by_contra, - exact (nat.totient_pos (zero_lt_iff.2 (ne.symm h))).ne hnm }, - { haveI := ne_zero.mk hzero, - rw [← map_cyclotomic_int _ R, ← map_cyclotomic_int _ R] at hnm, - replace hnm := map_injective (int.cast_ring_hom R) int.cast_injective hnm, - replace hnm := congr_arg (map (int.cast_ring_hom ℂ)) hnm, - rw [map_cyclotomic_int, map_cyclotomic_int] at hnm, - have hprim := complex.is_primitive_root_exp _ hzero, - have hroot := is_root_cyclotomic_iff.2 hprim, - rw hnm at hroot, - haveI hmzero : ne_zero m := ⟨λ h, by simpa [h] using hroot⟩, - rw is_root_cyclotomic_iff at hroot, - replace hprim := hprim.eq_order_of, - rwa [← is_primitive_root.eq_order_of hroot] at hprim} -end - lemma eq_cyclotomic_iff {R : Type*} [comm_ring R] {n : ℕ} (hpos: 0 < n) (P : R[X]) : P = cyclotomic n R ↔ P * (∏ i in nat.proper_divisors n, polynomial.cyclotomic i R) = X ^ n - 1 := @@ -778,235 +647,4 @@ end end order -section minpoly - -open is_primitive_root complex - -/-- The minimal polynomial of a primitive `n`-th root of unity `μ` divides `cyclotomic n ℤ`. -/ -lemma _root_.is_primitive_root.minpoly_dvd_cyclotomic {n : ℕ} {K : Type*} [field K] {μ : K} - (h : is_primitive_root μ n) (hpos : 0 < n) [char_zero K] : - minpoly ℤ μ ∣ cyclotomic n ℤ := -begin - apply minpoly.is_integrally_closed_dvd (is_integral h hpos), - simpa [aeval_def, eval₂_eq_eval_map, is_root.def] using is_root_cyclotomic hpos h -end - -lemma _root_.is_primitive_root.minpoly_eq_cyclotomic_of_irreducible {K : Type*} [field K] - {R : Type*} [comm_ring R] [is_domain R] {μ : R} {n : ℕ} [algebra K R] (hμ : is_primitive_root μ n) - (h : irreducible $ cyclotomic n K) [ne_zero (n : K)] : cyclotomic n K = minpoly K μ := -begin - haveI := ne_zero.of_no_zero_smul_divisors K R n, - refine minpoly.eq_of_irreducible_of_monic h _ (cyclotomic.monic n K), - rwa [aeval_def, eval₂_eq_eval_map, map_cyclotomic, ←is_root.def, is_root_cyclotomic_iff] -end - -/-- `cyclotomic n ℤ` is the minimal polynomial of a primitive `n`-th root of unity `μ`. -/ -lemma cyclotomic_eq_minpoly {n : ℕ} {K : Type*} [field K] {μ : K} - (h : is_primitive_root μ n) (hpos : 0 < n) [char_zero K] : - cyclotomic n ℤ = minpoly ℤ μ := -begin - refine eq_of_monic_of_dvd_of_nat_degree_le (minpoly.monic (is_integral h hpos)) - (cyclotomic.monic n ℤ) (h.minpoly_dvd_cyclotomic hpos) _, - simpa [nat_degree_cyclotomic n ℤ] using totient_le_degree_minpoly h -end - -/-- `cyclotomic n ℚ` is the minimal polynomial of a primitive `n`-th root of unity `μ`. -/ -lemma cyclotomic_eq_minpoly_rat {n : ℕ} {K : Type*} [field K] {μ : K} - (h : is_primitive_root μ n) (hpos : 0 < n) [char_zero K] : - cyclotomic n ℚ = minpoly ℚ μ := -begin - rw [← map_cyclotomic_int, cyclotomic_eq_minpoly h hpos], - exact (minpoly.is_integrally_closed_eq_field_fractions' _ (is_integral h hpos)).symm -end - -/-- `cyclotomic n ℤ` is irreducible. -/ -lemma cyclotomic.irreducible {n : ℕ} (hpos : 0 < n) : irreducible (cyclotomic n ℤ) := -begin - rw [cyclotomic_eq_minpoly (is_primitive_root_exp n hpos.ne') hpos], - apply minpoly.irreducible, - exact (is_primitive_root_exp n hpos.ne').is_integral hpos, -end - -/-- `cyclotomic n ℚ` is irreducible. -/ -lemma cyclotomic.irreducible_rat {n : ℕ} (hpos : 0 < n) : irreducible (cyclotomic n ℚ) := -begin - rw [← map_cyclotomic_int], - exact (is_primitive.irreducible_iff_irreducible_map_fraction_map (cyclotomic.is_primitive n ℤ)).1 - (cyclotomic.irreducible hpos), -end - -/-- If `n ≠ m`, then `(cyclotomic n ℚ)` and `(cyclotomic m ℚ)` are coprime. -/ -lemma cyclotomic.is_coprime_rat {n m : ℕ} (h : n ≠ m) : - is_coprime (cyclotomic n ℚ) (cyclotomic m ℚ) := -begin - rcases n.eq_zero_or_pos with rfl | hnzero, - { exact is_coprime_one_left }, - rcases m.eq_zero_or_pos with rfl | hmzero, - { exact is_coprime_one_right }, - rw (irreducible.coprime_iff_not_dvd $ cyclotomic.irreducible_rat $ hnzero), - exact (λ hdiv, h $ cyclotomic_injective $ eq_of_monic_of_associated (cyclotomic.monic n ℚ) - (cyclotomic.monic m ℚ) $ irreducible.associated_of_dvd (cyclotomic.irreducible_rat - hnzero) (cyclotomic.irreducible_rat hmzero) hdiv), -end - -end minpoly - -section expand - -/-- If `p` is a prime such that `¬ p ∣ n`, then -`expand R p (cyclotomic n R) = (cyclotomic (n * p) R) * (cyclotomic n R)`. -/ -@[simp] lemma cyclotomic_expand_eq_cyclotomic_mul {p n : ℕ} (hp : nat.prime p) (hdiv : ¬p ∣ n) - (R : Type*) [comm_ring R] : - expand R p (cyclotomic n R) = (cyclotomic (n * p) R) * (cyclotomic n R) := -begin - rcases nat.eq_zero_or_pos n with rfl | hnpos, - { simp }, - haveI := ne_zero.of_pos hnpos, - suffices : expand ℤ p (cyclotomic n ℤ) = (cyclotomic (n * p) ℤ) * (cyclotomic n ℤ), - { rw [← map_cyclotomic_int, ← map_expand, this, polynomial.map_mul, map_cyclotomic_int] }, - refine eq_of_monic_of_dvd_of_nat_degree_le ((cyclotomic.monic _ _).mul - (cyclotomic.monic _ _)) ((cyclotomic.monic n ℤ).expand hp.pos) _ _, - { refine (is_primitive.int.dvd_iff_map_cast_dvd_map_cast _ _ (is_primitive.mul - (cyclotomic.is_primitive (n * p) ℤ) (cyclotomic.is_primitive n ℤ)) - ((cyclotomic.monic n ℤ).expand hp.pos).is_primitive).2 _, - rw [polynomial.map_mul, map_cyclotomic_int, map_cyclotomic_int, map_expand, map_cyclotomic_int], - refine is_coprime.mul_dvd (cyclotomic.is_coprime_rat (λ h, _)) _ _, - { replace h : n * p = n * 1 := by simp [h], - exact nat.prime.ne_one hp (mul_left_cancel₀ hnpos.ne' h) }, - { have hpos : 0 < n * p := mul_pos hnpos hp.pos, - have hprim := complex.is_primitive_root_exp _ hpos.ne', - rw [cyclotomic_eq_minpoly_rat hprim hpos], - refine @minpoly.dvd ℚ ℂ _ _ algebra_rat _ _ _, - rw [aeval_def, ← eval_map, map_expand, map_cyclotomic, expand_eval, ← is_root.def, - is_root_cyclotomic_iff], - convert is_primitive_root.pow_of_dvd hprim hp.ne_zero (dvd_mul_left p n), - rw [nat.mul_div_cancel _ (nat.prime.pos hp)] }, - { have hprim := complex.is_primitive_root_exp _ hnpos.ne.symm, - rw [cyclotomic_eq_minpoly_rat hprim hnpos], - refine @minpoly.dvd ℚ ℂ _ _ algebra_rat _ _ _, - rw [aeval_def, ← eval_map, map_expand, expand_eval, ← is_root.def, - ← cyclotomic_eq_minpoly_rat hprim hnpos, map_cyclotomic, is_root_cyclotomic_iff], - exact is_primitive_root.pow_of_prime hprim hp hdiv,} }, - { rw [nat_degree_expand, nat_degree_cyclotomic, nat_degree_mul (cyclotomic_ne_zero _ ℤ) - (cyclotomic_ne_zero _ ℤ), nat_degree_cyclotomic, nat_degree_cyclotomic, mul_comm n, - nat.totient_mul ((nat.prime.coprime_iff_not_dvd hp).2 hdiv), - nat.totient_prime hp, mul_comm (p - 1), ← nat.mul_succ, nat.sub_one, - nat.succ_pred_eq_of_pos hp.pos] } -end - -/-- If `p` is a prime such that `p ∣ n`, then -`expand R p (cyclotomic n R) = cyclotomic (p * n) R`. -/ -@[simp] lemma cyclotomic_expand_eq_cyclotomic {p n : ℕ} (hp : nat.prime p) (hdiv : p ∣ n) - (R : Type*) [comm_ring R] : expand R p (cyclotomic n R) = cyclotomic (n * p) R := -begin - rcases n.eq_zero_or_pos with rfl | hzero, - { simp }, - haveI := ne_zero.of_pos hzero, - suffices : expand ℤ p (cyclotomic n ℤ) = cyclotomic (n * p) ℤ, - { rw [← map_cyclotomic_int, ← map_expand, this, map_cyclotomic_int] }, - refine eq_of_monic_of_dvd_of_nat_degree_le (cyclotomic.monic _ _) - ((cyclotomic.monic n ℤ).expand hp.pos) _ _, - { have hpos := nat.mul_pos hzero hp.pos, - have hprim := complex.is_primitive_root_exp _ hpos.ne.symm, - rw [cyclotomic_eq_minpoly hprim hpos], - refine minpoly.is_integrally_closed_dvd (hprim.is_integral hpos) _, - rw [aeval_def, ← eval_map, map_expand, map_cyclotomic, expand_eval, - ← is_root.def, is_root_cyclotomic_iff], - { convert is_primitive_root.pow_of_dvd hprim hp.ne_zero (dvd_mul_left p n), - rw [nat.mul_div_cancel _ hp.pos] } }, - { rw [nat_degree_expand, nat_degree_cyclotomic, nat_degree_cyclotomic, mul_comm n, - nat.totient_mul_of_prime_of_dvd hp hdiv, mul_comm] } -end - -/-- 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 - 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 - -/-- If `R` is of characteristic `p` and `¬p ∣ n`, then -`cyclotomic (n * p) R = (cyclotomic n R) ^ (p - 1)`. -/ -lemma cyclotomic_mul_prime_eq_pow_of_not_dvd (R : Type*) {p n : ℕ} [hp : fact (nat.prime p)] - [ring R] [char_p R p] (hn : ¬p ∣ n) : cyclotomic (n * p) R = (cyclotomic n R) ^ (p - 1) := -begin - suffices : cyclotomic (n * p) (zmod p) = (cyclotomic n (zmod p)) ^ (p - 1), - { rw [← map_cyclotomic _ (algebra_map (zmod p) R), ← map_cyclotomic _ (algebra_map (zmod p) R), - this, polynomial.map_pow] }, - apply mul_right_injective₀ (cyclotomic_ne_zero n $ zmod p), - rw [←pow_succ, tsub_add_cancel_of_le hp.out.one_lt.le, mul_comm, ← zmod.expand_card], - nth_rewrite 2 [← map_cyclotomic_int], - rw [← map_expand, cyclotomic_expand_eq_cyclotomic_mul hp.out hn, polynomial.map_mul, - map_cyclotomic, map_cyclotomic] -end - -/-- If `R` is of characteristic `p` and `p ∣ n`, then -`cyclotomic (n * p) R = (cyclotomic n R) ^ p`. -/ -lemma cyclotomic_mul_prime_dvd_eq_pow (R : Type*) {p n : ℕ} [hp : fact (nat.prime p)] [ring R] - [char_p R p] (hn : p ∣ n) : cyclotomic (n * p) R = (cyclotomic n R) ^ p := -begin - suffices : cyclotomic (n * p) (zmod p) = (cyclotomic n (zmod p)) ^ p, - { rw [← map_cyclotomic _ (algebra_map (zmod p) R), ← map_cyclotomic _ (algebra_map (zmod p) R), - this, polynomial.map_pow] }, - rw [← zmod.expand_card, ← map_cyclotomic_int n, ← map_expand, cyclotomic_expand_eq_cyclotomic - hp.out hn, map_cyclotomic, mul_comm] -end - -/-- If `R` is of characteristic `p` and `¬p ∣ m`, then -`cyclotomic (p ^ k * m) R = (cyclotomic m R) ^ (p ^ k - p ^ (k - 1))`. -/ -lemma cyclotomic_mul_prime_pow_eq (R : Type*) {p m : ℕ} [fact (nat.prime p)] - [ring R] [char_p R p] (hm : ¬p ∣ m) : - ∀ {k}, 0 < k → cyclotomic (p ^ k * m) R = (cyclotomic m R) ^ (p ^ k - p ^ (k - 1)) -| 1 _ := by rw [pow_one, nat.sub_self, pow_zero, mul_comm, - cyclotomic_mul_prime_eq_pow_of_not_dvd R hm] -| (a + 2) _ := -begin - have hdiv : p ∣ p ^ a.succ * m := ⟨p ^ a * m, by rw [← mul_assoc, pow_succ]⟩, - rw [pow_succ, mul_assoc, mul_comm, cyclotomic_mul_prime_dvd_eq_pow R hdiv, - cyclotomic_mul_prime_pow_eq a.succ_pos, ← pow_mul], - congr' 1, - simp only [tsub_zero, nat.succ_sub_succ_eq_sub], - rw [nat.mul_sub_right_distrib, mul_comm, pow_succ'] -end - -/-- If `R` is of characteristic `p` and `¬p ∣ m`, then `ζ` is a root of `cyclotomic (p ^ k * m) R` - if and only if it is a primitive `m`-th root of unity. -/ -lemma is_root_cyclotomic_prime_pow_mul_iff_of_char_p {m k p : ℕ} {R : Type*} [comm_ring R] - [is_domain R] [hp : fact (nat.prime p)] [hchar : char_p R p] {μ : R} [ne_zero (m : R)] : - (polynomial.cyclotomic (p ^ k * m) R).is_root μ ↔ is_primitive_root μ m := -begin - rcases k.eq_zero_or_pos with rfl | hk, - { rw [pow_zero, one_mul, is_root_cyclotomic_iff] }, - refine ⟨λ h, _, λ h, _⟩, - { rw [is_root.def, cyclotomic_mul_prime_pow_eq R (ne_zero.not_char_dvd R p m) hk, eval_pow] at h, - replace h := pow_eq_zero h, - rwa [← is_root.def, is_root_cyclotomic_iff] at h }, - { rw [← is_root_cyclotomic_iff, is_root.def] at h, - rw [cyclotomic_mul_prime_pow_eq R (ne_zero.not_char_dvd R p m) hk, - is_root.def, eval_pow, h, zero_pow], - simp only [tsub_pos_iff_lt], - apply pow_strict_mono_right hp.out.one_lt (nat.pred_lt hk.ne') } -end - -end char_p - end polynomial diff --git a/src/ring_theory/polynomial/cyclotomic/eval.lean b/src/ring_theory/polynomial/cyclotomic/eval.lean index 77a8f285d4c51..84bb6f3300a52 100644 --- a/src/ring_theory/polynomial/cyclotomic/eval.lean +++ b/src/ring_theory/polynomial/cyclotomic/eval.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Eric Rodriguez -/ -import ring_theory.polynomial.cyclotomic.basic +import ring_theory.polynomial.cyclotomic.roots import tactic.by_contra import topology.algebra.polynomial import number_theory.padics.padic_val diff --git a/src/ring_theory/polynomial/cyclotomic/expand.lean b/src/ring_theory/polynomial/cyclotomic/expand.lean new file mode 100644 index 0000000000000..83ae1209b818f --- /dev/null +++ b/src/ring_theory/polynomial/cyclotomic/expand.lean @@ -0,0 +1,186 @@ +/- +Copyright (c) 2020 Riccardo Brasca. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Riccardo Brasca +-/ + +import ring_theory.polynomial.cyclotomic.roots + +/-! +# Cyclotomic polynomials and `expand`. + +We gather results relating cyclotomic polynomials and `expand`. + +## Main results + +* `polynomial.cyclotomic_expand_eq_cyclotomic_mul` : If `p` is a prime such that `¬ p ∣ n`, then + `expand R p (cyclotomic n R) = (cyclotomic (n * p) R) * (cyclotomic n R)`. +* `polynomial.cyclotomic_expand_eq_cyclotomic` : If `p` is a prime such that `p ∣ n`, then + `expand R p (cyclotomic n R) = cyclotomic (p * n) R`. +* `polynomial.cyclotomic_mul_prime_eq_pow_of_not_dvd` : If `R` is of characteristic `p` and + `¬p ∣ n`, then `cyclotomic (n * p) R = (cyclotomic n R) ^ (p - 1)`. +* `polynomial.cyclotomic_mul_prime_dvd_eq_pow` : If `R` is of characteristic `p` and `p ∣ n`, then + `cyclotomic (n * p) R = (cyclotomic n R) ^ p`. +* `polynomial.cyclotomic_mul_prime_pow_eq` : If `R` is of characteristic `p` and `¬p ∣ m`, then + `cyclotomic (p ^ k * m) R = (cyclotomic m R) ^ (p ^ k - p ^ (k - 1))`. +* `polynomial.cyclotomic_mul_prime_pow_eq` : If `R` is of characteristic `p` and `¬p ∣ m`, then + `cyclotomic (p ^ k * m) R = (cyclotomic m R) ^ (p ^ k - p ^ (k - 1))`. +-/ + +namespace polynomial + +/-- If `p` is a prime such that `¬ p ∣ n`, then +`expand R p (cyclotomic n R) = (cyclotomic (n * p) R) * (cyclotomic n R)`. -/ +@[simp] lemma cyclotomic_expand_eq_cyclotomic_mul {p n : ℕ} (hp : nat.prime p) (hdiv : ¬p ∣ n) + (R : Type*) [comm_ring R] : + expand R p (cyclotomic n R) = (cyclotomic (n * p) R) * (cyclotomic n R) := +begin + rcases nat.eq_zero_or_pos n with rfl | hnpos, + { simp }, + haveI := ne_zero.of_pos hnpos, + suffices : expand ℤ p (cyclotomic n ℤ) = (cyclotomic (n * p) ℤ) * (cyclotomic n ℤ), + { rw [← map_cyclotomic_int, ← map_expand, this, polynomial.map_mul, map_cyclotomic_int] }, + refine eq_of_monic_of_dvd_of_nat_degree_le ((cyclotomic.monic _ _).mul + (cyclotomic.monic _ _)) ((cyclotomic.monic n ℤ).expand hp.pos) _ _, + { refine (is_primitive.int.dvd_iff_map_cast_dvd_map_cast _ _ (is_primitive.mul + (cyclotomic.is_primitive (n * p) ℤ) (cyclotomic.is_primitive n ℤ)) + ((cyclotomic.monic n ℤ).expand hp.pos).is_primitive).2 _, + rw [polynomial.map_mul, map_cyclotomic_int, map_cyclotomic_int, map_expand, map_cyclotomic_int], + refine is_coprime.mul_dvd (cyclotomic.is_coprime_rat (λ h, _)) _ _, + { replace h : n * p = n * 1 := by simp [h], + exact nat.prime.ne_one hp (mul_left_cancel₀ hnpos.ne' h) }, + { have hpos : 0 < n * p := mul_pos hnpos hp.pos, + have hprim := complex.is_primitive_root_exp _ hpos.ne', + rw [cyclotomic_eq_minpoly_rat hprim hpos], + refine @minpoly.dvd ℚ ℂ _ _ algebra_rat _ _ _, + rw [aeval_def, ← eval_map, map_expand, map_cyclotomic, expand_eval, ← is_root.def, + is_root_cyclotomic_iff], + convert is_primitive_root.pow_of_dvd hprim hp.ne_zero (dvd_mul_left p n), + rw [nat.mul_div_cancel _ (nat.prime.pos hp)] }, + { have hprim := complex.is_primitive_root_exp _ hnpos.ne.symm, + rw [cyclotomic_eq_minpoly_rat hprim hnpos], + refine @minpoly.dvd ℚ ℂ _ _ algebra_rat _ _ _, + rw [aeval_def, ← eval_map, map_expand, expand_eval, ← is_root.def, + ← cyclotomic_eq_minpoly_rat hprim hnpos, map_cyclotomic, is_root_cyclotomic_iff], + exact is_primitive_root.pow_of_prime hprim hp hdiv,} }, + { rw [nat_degree_expand, nat_degree_cyclotomic, nat_degree_mul (cyclotomic_ne_zero _ ℤ) + (cyclotomic_ne_zero _ ℤ), nat_degree_cyclotomic, nat_degree_cyclotomic, mul_comm n, + nat.totient_mul ((nat.prime.coprime_iff_not_dvd hp).2 hdiv), + nat.totient_prime hp, mul_comm (p - 1), ← nat.mul_succ, nat.sub_one, + nat.succ_pred_eq_of_pos hp.pos] } +end + +/-- If `p` is a prime such that `p ∣ n`, then +`expand R p (cyclotomic n R) = cyclotomic (p * n) R`. -/ +@[simp] lemma cyclotomic_expand_eq_cyclotomic {p n : ℕ} (hp : nat.prime p) (hdiv : p ∣ n) + (R : Type*) [comm_ring R] : expand R p (cyclotomic n R) = cyclotomic (n * p) R := +begin + rcases n.eq_zero_or_pos with rfl | hzero, + { simp }, + haveI := ne_zero.of_pos hzero, + suffices : expand ℤ p (cyclotomic n ℤ) = cyclotomic (n * p) ℤ, + { rw [← map_cyclotomic_int, ← map_expand, this, map_cyclotomic_int] }, + refine eq_of_monic_of_dvd_of_nat_degree_le (cyclotomic.monic _ _) + ((cyclotomic.monic n ℤ).expand hp.pos) _ _, + { have hpos := nat.mul_pos hzero hp.pos, + have hprim := complex.is_primitive_root_exp _ hpos.ne.symm, + rw [cyclotomic_eq_minpoly hprim hpos], + refine minpoly.is_integrally_closed_dvd (hprim.is_integral hpos) _, + rw [aeval_def, ← eval_map, map_expand, map_cyclotomic, expand_eval, + ← is_root.def, is_root_cyclotomic_iff], + { convert is_primitive_root.pow_of_dvd hprim hp.ne_zero (dvd_mul_left p n), + rw [nat.mul_div_cancel _ hp.pos] } }, + { rw [nat_degree_expand, nat_degree_cyclotomic, nat_degree_cyclotomic, mul_comm n, + nat.totient_mul_of_prime_of_dvd hp hdiv, mul_comm] } +end + +/-- 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 + 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 + +section char_p + +/-- If `R` is of characteristic `p` and `¬p ∣ n`, then +`cyclotomic (n * p) R = (cyclotomic n R) ^ (p - 1)`. -/ +lemma cyclotomic_mul_prime_eq_pow_of_not_dvd (R : Type*) {p n : ℕ} [hp : fact (nat.prime p)] + [ring R] [char_p R p] (hn : ¬p ∣ n) : cyclotomic (n * p) R = (cyclotomic n R) ^ (p - 1) := +begin + suffices : cyclotomic (n * p) (zmod p) = (cyclotomic n (zmod p)) ^ (p - 1), + { rw [← map_cyclotomic _ (algebra_map (zmod p) R), ← map_cyclotomic _ (algebra_map (zmod p) R), + this, polynomial.map_pow] }, + apply mul_right_injective₀ (cyclotomic_ne_zero n $ zmod p), + rw [←pow_succ, tsub_add_cancel_of_le hp.out.one_lt.le, mul_comm, ← zmod.expand_card], + nth_rewrite 2 [← map_cyclotomic_int], + rw [← map_expand, cyclotomic_expand_eq_cyclotomic_mul hp.out hn, polynomial.map_mul, + map_cyclotomic, map_cyclotomic] +end + +/-- If `R` is of characteristic `p` and `p ∣ n`, then +`cyclotomic (n * p) R = (cyclotomic n R) ^ p`. -/ +lemma cyclotomic_mul_prime_dvd_eq_pow (R : Type*) {p n : ℕ} [hp : fact (nat.prime p)] [ring R] + [char_p R p] (hn : p ∣ n) : cyclotomic (n * p) R = (cyclotomic n R) ^ p := +begin + suffices : cyclotomic (n * p) (zmod p) = (cyclotomic n (zmod p)) ^ p, + { rw [← map_cyclotomic _ (algebra_map (zmod p) R), ← map_cyclotomic _ (algebra_map (zmod p) R), + this, polynomial.map_pow] }, + rw [← zmod.expand_card, ← map_cyclotomic_int n, ← map_expand, cyclotomic_expand_eq_cyclotomic + hp.out hn, map_cyclotomic, mul_comm] +end + +/-- If `R` is of characteristic `p` and `¬p ∣ m`, then +`cyclotomic (p ^ k * m) R = (cyclotomic m R) ^ (p ^ k - p ^ (k - 1))`. -/ +lemma cyclotomic_mul_prime_pow_eq (R : Type*) {p m : ℕ} [fact (nat.prime p)] + [ring R] [char_p R p] (hm : ¬p ∣ m) : + ∀ {k}, 0 < k → cyclotomic (p ^ k * m) R = (cyclotomic m R) ^ (p ^ k - p ^ (k - 1)) +| 1 _ := by rw [pow_one, nat.sub_self, pow_zero, mul_comm, + cyclotomic_mul_prime_eq_pow_of_not_dvd R hm] +| (a + 2) _ := +begin + have hdiv : p ∣ p ^ a.succ * m := ⟨p ^ a * m, by rw [← mul_assoc, pow_succ]⟩, + rw [pow_succ, mul_assoc, mul_comm, cyclotomic_mul_prime_dvd_eq_pow R hdiv, + cyclotomic_mul_prime_pow_eq a.succ_pos, ← pow_mul], + congr' 1, + simp only [tsub_zero, nat.succ_sub_succ_eq_sub], + rw [nat.mul_sub_right_distrib, mul_comm, pow_succ'] +end + +/-- If `R` is of characteristic `p` and `¬p ∣ m`, then `ζ` is a root of `cyclotomic (p ^ k * m) R` + if and only if it is a primitive `m`-th root of unity. -/ +lemma is_root_cyclotomic_prime_pow_mul_iff_of_char_p {m k p : ℕ} {R : Type*} [comm_ring R] + [is_domain R] [hp : fact (nat.prime p)] [hchar : char_p R p] {μ : R} [ne_zero (m : R)] : + (polynomial.cyclotomic (p ^ k * m) R).is_root μ ↔ is_primitive_root μ m := +begin + rcases k.eq_zero_or_pos with rfl | hk, + { rw [pow_zero, one_mul, is_root_cyclotomic_iff] }, + refine ⟨λ h, _, λ h, _⟩, + { rw [is_root.def, cyclotomic_mul_prime_pow_eq R (ne_zero.not_char_dvd R p m) hk, eval_pow] at h, + replace h := pow_eq_zero h, + rwa [← is_root.def, is_root_cyclotomic_iff] at h }, + { rw [← is_root_cyclotomic_iff, is_root.def] at h, + rw [cyclotomic_mul_prime_pow_eq R (ne_zero.not_char_dvd R p m) hk, + is_root.def, eval_pow, h, zero_pow], + simp only [tsub_pos_iff_lt], + apply pow_strict_mono_right hp.out.one_lt (nat.pred_lt hk.ne') } +end + +end char_p + +end polynomial diff --git a/src/ring_theory/polynomial/cyclotomic/roots.lean b/src/ring_theory/polynomial/cyclotomic/roots.lean new file mode 100644 index 0000000000000..78bdafb88dbc6 --- /dev/null +++ b/src/ring_theory/polynomial/cyclotomic/roots.lean @@ -0,0 +1,243 @@ +/- +Copyright (c) 2020 Riccardo Brasca. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Riccardo Brasca +-/ + +import ring_theory.polynomial.cyclotomic.basic + +/-! +# Roots of cyclotomic polynomials. + +We gather results about roots of cyclotomic polynomials. In particular we show in +`polynomial.cyclotomic_eq_minpoly` that `cyclotomic n R` is the minimal polynomial of a primitive +root of unity. + +## Main results + +* `is_primitive_root.is_root_cyclotomic` : Any `n`-th primitive root of unity is a root of + `cyclotomic n R`. +* `is_root_cyclotomic_iff` : if `ne_zero (n : R)`, then `μ` is a root of `cyclotomic n R` + if and only if `μ` is a primitive root of unity. +* `polynomial.cyclotomic_eq_minpoly` : `cyclotomic n ℤ` is the minimal polynomial of a primitive + `n`-th root of unity `μ`. +* `polynomial.cyclotomic.irreducible` : `cyclotomic n ℤ` is irreducible. + +## Implementation details + +To prove `polynomial.cyclotomic.irreducible`, the irreducibility of `cyclotomic n ℤ`, we show in +`polynomial.cyclotomic_eq_minpoly` that `cyclotomic n ℤ` is the minimal polynomial of any `n`-th +primitive root of unity `μ : K`, where `K` is a field of characteristic `0`. +-/ + +open_locale big_operators + +namespace polynomial + +variables {R : Type*} [comm_ring R] {n : ℕ} + +lemma is_root_of_unity_of_root_cyclotomic {ζ : R} {i : ℕ} + (hi : i ∈ n.divisors) (h : (cyclotomic i R).is_root ζ) : ζ ^ n = 1 := +begin + rcases n.eq_zero_or_pos with rfl | hn, + { exact pow_zero _ }, + have := congr_arg (eval ζ) (prod_cyclotomic_eq_X_pow_sub_one hn R).symm, + rw [eval_sub, eval_pow, eval_X, eval_one] at this, + convert eq_add_of_sub_eq' this, + convert (add_zero _).symm, + apply eval_eq_zero_of_dvd_of_eval_eq_zero _ h, + exact finset.dvd_prod_of_mem _ hi +end + +section is_domain + +variable [is_domain R] + +lemma _root_.is_root_of_unity_iff (h : 0 < n) (R : Type*) [comm_ring R] [is_domain R] + {ζ : R} : ζ ^ n = 1 ↔ ∃ i ∈ n.divisors, (cyclotomic i R).is_root ζ := +by rw [←mem_nth_roots h, nth_roots, mem_roots $ X_pow_sub_C_ne_zero h _, + C_1, ←prod_cyclotomic_eq_X_pow_sub_one h, is_root_prod]; apply_instance + +/-- Any `n`-th primitive root of unity is a root of `cyclotomic n R`.-/ +lemma _root_.is_primitive_root.is_root_cyclotomic (hpos : 0 < n) {μ : R} + (h : is_primitive_root μ n) : is_root (cyclotomic n R) μ := +begin + rw [← mem_roots (cyclotomic_ne_zero n R), + cyclotomic_eq_prod_X_sub_primitive_roots h, roots_prod_X_sub_C, ← finset.mem_def], + rwa [← mem_primitive_roots hpos] at h, +end + +private lemma is_root_cyclotomic_iff' {n : ℕ} {K : Type*} [field K] {μ : K} [ne_zero (n : K)] : + is_root (cyclotomic n K) μ ↔ is_primitive_root μ n := +begin + -- in this proof, `o` stands for `order_of μ` + have hnpos : 0 < n := (ne_zero.of_ne_zero_coe K).out.bot_lt, + refine ⟨λ hμ, _, is_primitive_root.is_root_cyclotomic hnpos⟩, + have hμn : μ ^ n = 1, + { rw is_root_of_unity_iff hnpos _, + exact ⟨n, n.mem_divisors_self hnpos.ne', hμ⟩, + all_goals { apply_instance } }, + by_contra hnμ, + have ho : 0 < order_of μ, + { apply order_of_pos', + rw is_of_fin_order_iff_pow_eq_one, + exact ⟨n, hnpos, hμn⟩ }, + have := pow_order_of_eq_one μ, + rw is_root_of_unity_iff ho at this, + obtain ⟨i, hio, hiμ⟩ := this, + replace hio := nat.dvd_of_mem_divisors hio, + rw is_primitive_root.not_iff at hnμ, + rw ←order_of_dvd_iff_pow_eq_one at hμn, + have key : i < n := (nat.le_of_dvd ho hio).trans_lt ((nat.le_of_dvd hnpos hμn).lt_of_ne hnμ), + have key' : i ∣ n := hio.trans hμn, + rw ←polynomial.dvd_iff_is_root at hμ hiμ, + have hni : {i, n} ⊆ n.divisors, + { simpa [finset.insert_subset, key'] using hnpos.ne' }, + obtain ⟨k, hk⟩ := hiμ, + obtain ⟨j, hj⟩ := hμ, + have := prod_cyclotomic_eq_X_pow_sub_one hnpos K, + rw [←finset.prod_sdiff hni, finset.prod_pair key.ne, hk, hj] at this, + have hn := (X_pow_sub_one_separable_iff.mpr $ ne_zero.nat_cast_ne n K).squarefree, + rw [←this, squarefree] at hn, + contrapose! hn, + refine ⟨X - C μ, ⟨(∏ x in n.divisors \ {i, n}, cyclotomic x K) * k * j, by ring⟩, _⟩, + simp [polynomial.is_unit_iff_degree_eq_zero], + all_goals { apply_instance } +end + +lemma is_root_cyclotomic_iff [ne_zero (n : R)] {μ : R} : + is_root (cyclotomic n R) μ ↔ is_primitive_root μ n := +begin + have hf : function.injective _ := is_fraction_ring.injective R (fraction_ring R), + haveI : ne_zero (n : fraction_ring R) := ne_zero.nat_of_injective hf, + rw [←is_root_map_iff hf, ←is_primitive_root.map_iff_of_injective hf, map_cyclotomic, + ←is_root_cyclotomic_iff'] +end + +lemma roots_cyclotomic_nodup [ne_zero (n : R)] : (cyclotomic n R).roots.nodup := +begin + obtain h | ⟨ζ, hζ⟩ := (cyclotomic n R).roots.empty_or_exists_mem, + { exact h.symm ▸ multiset.nodup_zero }, + rw [mem_roots $ cyclotomic_ne_zero n R, is_root_cyclotomic_iff] at hζ, + refine multiset.nodup_of_le (roots.le_of_dvd (X_pow_sub_C_ne_zero + (ne_zero.pos_of_ne_zero_coe R) 1) $ cyclotomic.dvd_X_pow_sub_one n R) hζ.nth_roots_nodup, +end + +lemma cyclotomic.roots_to_finset_eq_primitive_roots [ne_zero (n : R)] : + (⟨(cyclotomic n R).roots, roots_cyclotomic_nodup⟩ : finset _) = primitive_roots n R := +by { ext, simp [cyclotomic_ne_zero n R, is_root_cyclotomic_iff, + mem_primitive_roots, ne_zero.pos_of_ne_zero_coe R] } + +lemma cyclotomic.roots_eq_primitive_roots_val [ne_zero (n : R)] : + (cyclotomic n R).roots = (primitive_roots n R).val := +by rw ←cyclotomic.roots_to_finset_eq_primitive_roots + +/-- If `R` is of characteristic zero, then `ζ` is a root of `cyclotomic n R` if and only if it is a +primitive `n`-th root of unity. -/ +lemma is_root_cyclotomic_iff_char_zero {n : ℕ} {R : Type*} [comm_ring R] [is_domain R] + [char_zero R] {μ : R} (hn : 0 < n) : + (polynomial.cyclotomic n R).is_root μ ↔ is_primitive_root μ n := +by { letI := ne_zero.of_gt hn, exact is_root_cyclotomic_iff } + +end is_domain + +/-- Over a ring `R` of characteristic zero, `λ n, cyclotomic n R` is injective. -/ +lemma cyclotomic_injective [char_zero R] : + function.injective (λ n, cyclotomic n R) := +begin + intros n m hnm, + simp only at hnm, + rcases eq_or_ne n 0 with rfl | hzero, + { rw [cyclotomic_zero] at hnm, + replace hnm := congr_arg nat_degree hnm, + rw [nat_degree_one, nat_degree_cyclotomic] at hnm, + by_contra, + exact (nat.totient_pos (zero_lt_iff.2 (ne.symm h))).ne hnm }, + { haveI := ne_zero.mk hzero, + rw [← map_cyclotomic_int _ R, ← map_cyclotomic_int _ R] at hnm, + replace hnm := map_injective (int.cast_ring_hom R) int.cast_injective hnm, + replace hnm := congr_arg (map (int.cast_ring_hom ℂ)) hnm, + rw [map_cyclotomic_int, map_cyclotomic_int] at hnm, + have hprim := complex.is_primitive_root_exp _ hzero, + have hroot := is_root_cyclotomic_iff.2 hprim, + rw hnm at hroot, + haveI hmzero : ne_zero m := ⟨λ h, by simpa [h] using hroot⟩, + rw is_root_cyclotomic_iff at hroot, + replace hprim := hprim.eq_order_of, + rwa [← is_primitive_root.eq_order_of hroot] at hprim} +end + +/-- The minimal polynomial of a primitive `n`-th root of unity `μ` divides `cyclotomic n ℤ`. -/ +lemma _root_.is_primitive_root.minpoly_dvd_cyclotomic {n : ℕ} {K : Type*} [field K] {μ : K} + (h : is_primitive_root μ n) (hpos : 0 < n) [char_zero K] : + minpoly ℤ μ ∣ cyclotomic n ℤ := +begin + apply minpoly.is_integrally_closed_dvd (h.is_integral hpos), + simpa [aeval_def, eval₂_eq_eval_map, is_root.def] using h.is_root_cyclotomic hpos +end + +section minpoly + +open is_primitive_root complex + +lemma _root_.is_primitive_root.minpoly_eq_cyclotomic_of_irreducible {K : Type*} [field K] + {R : Type*} [comm_ring R] [is_domain R] {μ : R} {n : ℕ} [algebra K R] (hμ : is_primitive_root μ n) + (h : irreducible $ cyclotomic n K) [ne_zero (n : K)] : cyclotomic n K = minpoly K μ := +begin + haveI := ne_zero.of_no_zero_smul_divisors K R n, + refine minpoly.eq_of_irreducible_of_monic h _ (cyclotomic.monic n K), + rwa [aeval_def, eval₂_eq_eval_map, map_cyclotomic, ←is_root.def, is_root_cyclotomic_iff] +end + +/-- `cyclotomic n ℤ` is the minimal polynomial of a primitive `n`-th root of unity `μ`. -/ +lemma cyclotomic_eq_minpoly {n : ℕ} {K : Type*} [field K] {μ : K} + (h : is_primitive_root μ n) (hpos : 0 < n) [char_zero K] : + cyclotomic n ℤ = minpoly ℤ μ := +begin + refine eq_of_monic_of_dvd_of_nat_degree_le (minpoly.monic (is_integral h hpos)) + (cyclotomic.monic n ℤ) (h.minpoly_dvd_cyclotomic hpos) _, + simpa [nat_degree_cyclotomic n ℤ] using totient_le_degree_minpoly h +end + +/-- `cyclotomic n ℚ` is the minimal polynomial of a primitive `n`-th root of unity `μ`. -/ +lemma cyclotomic_eq_minpoly_rat {n : ℕ} {K : Type*} [field K] {μ : K} + (h : is_primitive_root μ n) (hpos : 0 < n) [char_zero K] : + cyclotomic n ℚ = minpoly ℚ μ := +begin + rw [← map_cyclotomic_int, cyclotomic_eq_minpoly h hpos], + exact (minpoly.is_integrally_closed_eq_field_fractions' _ (is_integral h hpos)).symm +end + +/-- `cyclotomic n ℤ` is irreducible. -/ +lemma cyclotomic.irreducible {n : ℕ} (hpos : 0 < n) : irreducible (cyclotomic n ℤ) := +begin + rw [cyclotomic_eq_minpoly (is_primitive_root_exp n hpos.ne') hpos], + apply minpoly.irreducible, + exact (is_primitive_root_exp n hpos.ne').is_integral hpos, +end + +/-- `cyclotomic n ℚ` is irreducible. -/ +lemma cyclotomic.irreducible_rat {n : ℕ} (hpos : 0 < n) : irreducible (cyclotomic n ℚ) := +begin + rw [← map_cyclotomic_int], + exact (is_primitive.irreducible_iff_irreducible_map_fraction_map (cyclotomic.is_primitive n ℤ)).1 + (cyclotomic.irreducible hpos), +end + +/-- If `n ≠ m`, then `(cyclotomic n ℚ)` and `(cyclotomic m ℚ)` are coprime. -/ +lemma cyclotomic.is_coprime_rat {n m : ℕ} (h : n ≠ m) : + is_coprime (cyclotomic n ℚ) (cyclotomic m ℚ) := +begin + rcases n.eq_zero_or_pos with rfl | hnzero, + { exact is_coprime_one_left }, + rcases m.eq_zero_or_pos with rfl | hmzero, + { exact is_coprime_one_right }, + rw (irreducible.coprime_iff_not_dvd $ cyclotomic.irreducible_rat $ hnzero), + exact (λ hdiv, h $ cyclotomic_injective $ eq_of_monic_of_associated (cyclotomic.monic n ℚ) + (cyclotomic.monic m ℚ) $ irreducible.associated_of_dvd (cyclotomic.irreducible_rat + hnzero) (cyclotomic.irreducible_rat hmzero) hdiv), +end + +end minpoly + +end polynomial diff --git a/src/ring_theory/polynomial/eisenstein/is_integral.lean b/src/ring_theory/polynomial/eisenstein/is_integral.lean index 8c44b9246e35d..0704e82c63a77 100644 --- a/src/ring_theory/polynomial/eisenstein/is_integral.lean +++ b/src/ring_theory/polynomial/eisenstein/is_integral.lean @@ -6,7 +6,7 @@ Authors: Riccardo Brasca import data.nat.choose.dvd import ring_theory.integrally_closed import ring_theory.norm -import ring_theory.polynomial.cyclotomic.basic +import ring_theory.polynomial.cyclotomic.expand /-! # Eisenstein polynomials From 0372d31fb681ef40a687506bc5870fd55ebc8bb9 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 24 May 2023 15:42:16 +0000 Subject: [PATCH 1071/1291] feat(algebra/group/opposite): `is_unit` lemmas for `mul_opposite` (#19080) I need these to build lemmas about right-multiplication by a unit out of lemmas about the left-action by a unit of the opposite ring, and this lets me convert that hypothesis. --- src/algebra/group/opposite.lean | 15 +++++++++++++++ 1 file changed, 15 insertions(+) diff --git a/src/algebra/group/opposite.lean b/src/algebra/group/opposite.lean index 18f3d07d6c75f..11891aa81b415 100644 --- a/src/algebra/group/opposite.lean +++ b/src/algebra/group/opposite.lean @@ -347,6 +347,21 @@ lemma units.coe_op_equiv_symm {M} [monoid M] (u : (Mˣ)ᵐᵒᵖ) : (units.op_equiv.symm u : Mᵐᵒᵖ) = op (u.unop : M) := rfl +@[to_additive] +lemma is_unit.op {M} [monoid M] {m : M} (h : is_unit m) : is_unit (op m) := +let ⟨u, hu⟩ := h in hu ▸ ⟨units.op_equiv.symm (op u), rfl⟩ + +@[to_additive] +lemma is_unit.unop {M} [monoid M] {m : Mᵐᵒᵖ} (h : is_unit m) : is_unit (unop m) := +let ⟨u, hu⟩ := h in hu ▸ ⟨unop (units.op_equiv u), rfl⟩ + +@[simp, to_additive] +lemma is_unit_op {M} [monoid M] {m : M} : is_unit (op m) ↔ is_unit m := ⟨is_unit.unop, is_unit.op⟩ + +@[simp, to_additive] +lemma is_unit_unop {M} [monoid M] {m : Mᵐᵒᵖ} : is_unit (unop m) ↔ is_unit m := +⟨is_unit.op, is_unit.unop⟩ + /-- 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 From ef95945cd48c932c9e034872bd25c3c220d9c946 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 24 May 2023 15:42:17 +0000 Subject: [PATCH 1072/1291] =?UTF-8?q?feat(measure=5Ftheory/function/strong?= =?UTF-8?q?ly=5Fmeasurable/basic):=20generalize=20to=20`is=5Funit=20c`=20f?= =?UTF-8?q?rom=20`c=20=E2=89=A0=200`=20(#19081)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We already have this generalization for `measurable_const_smul_iff` and `ae_measurable_const_smul_iff`. --- .../function/strongly_measurable/basic.lean | 37 +++++++++---------- 1 file changed, 18 insertions(+), 19 deletions(-) diff --git a/src/measure_theory/function/strongly_measurable/basic.lean b/src/measure_theory/function/strongly_measurable/basic.lean index 6315583d2588b..c7220c3cb2d68 100644 --- a/src/measure_theory/function/strongly_measurable/basic.lean +++ b/src/measure_theory/function/strongly_measurable/basic.lean @@ -445,24 +445,24 @@ continuous_smul.comp_strongly_measurable (hf.prod_mk strongly_measurable_const) end arithmetic section mul_action - -variables [topological_space β] {G : Type*} [group G] [mul_action G β] - [has_continuous_const_smul G β] +variables {M G G₀ : Type*} +variables [topological_space β] +variables [monoid M] [mul_action M β] [has_continuous_const_smul M β] +variables [group G] [mul_action G β] [has_continuous_const_smul G β] +variables [group_with_zero G₀] [mul_action G₀ β] [has_continuous_const_smul G₀ β] lemma _root_.strongly_measurable_const_smul_iff {m : measurable_space α} (c : G) : strongly_measurable (λ x, c • f x) ↔ strongly_measurable f := ⟨λ h, by simpa only [inv_smul_smul] using h.const_smul' c⁻¹, λ h, h.const_smul c⟩ -variables {G₀ : Type*} [group_with_zero G₀] [mul_action G₀ β] - [has_continuous_const_smul G₀ β] +lemma _root_.is_unit.strongly_measurable_const_smul_iff {m : measurable_space α} {c : M} + (hc : is_unit c) : + strongly_measurable (λ x, c • f x) ↔ strongly_measurable f := +let ⟨u, hu⟩ := hc in hu ▸ strongly_measurable_const_smul_iff u lemma _root_.strongly_measurable_const_smul_iff₀ {m : measurable_space α} {c : G₀} (hc : c ≠ 0) : strongly_measurable (λ x, c • f x) ↔ strongly_measurable f := -begin - refine ⟨λ h, _, λ h, h.const_smul c⟩, - convert h.const_smul' c⁻¹, - simp [smul_smul, inv_mul_cancel hc] -end +(is_unit.mk0 _ hc).strongly_measurable_const_smul_iff end mul_action @@ -1667,23 +1667,22 @@ end normed_space section mul_action -variables {G : Type*} [group G] [mul_action G β] - [has_continuous_const_smul G β] +variables {M G G₀ : Type*} +variables [monoid M] [mul_action M β] [has_continuous_const_smul M β] +variables [group G] [mul_action G β] [has_continuous_const_smul G β] +variables [group_with_zero G₀] [mul_action G₀ β] [has_continuous_const_smul G₀ β] lemma _root_.ae_strongly_measurable_const_smul_iff (c : G) : ae_strongly_measurable (λ x, c • f x) μ ↔ ae_strongly_measurable f μ := ⟨λ h, by simpa only [inv_smul_smul] using h.const_smul' c⁻¹, λ h, h.const_smul c⟩ -variables {G₀ : Type*} [group_with_zero G₀] [mul_action G₀ β] - [has_continuous_const_smul G₀ β] +lemma _root_.is_unit.ae_strongly_measurable_const_smul_iff {c : M} (hc : is_unit c) : + ae_strongly_measurable (λ x, c • f x) μ ↔ ae_strongly_measurable f μ := +let ⟨u, hu⟩ := hc in hu ▸ ae_strongly_measurable_const_smul_iff u lemma _root_.ae_strongly_measurable_const_smul_iff₀ {c : G₀} (hc : c ≠ 0) : ae_strongly_measurable (λ x, c • f x) μ ↔ ae_strongly_measurable f μ := -begin - refine ⟨λ h, _, λ h, h.const_smul c⟩, - convert h.const_smul' c⁻¹, - simp [smul_smul, inv_mul_cancel hc] -end +(is_unit.mk0 _ hc).ae_strongly_measurable_const_smul_iff end mul_action From 96d2ccbe330472f04749ae707251856775dddee5 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 24 May 2023 21:28:29 +0000 Subject: [PATCH 1073/1291] fix: bump to Lean 3.51.1 (#19088) This should fix the web editor --- leanpkg.toml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/leanpkg.toml b/leanpkg.toml index 06b84b6c77b86..fe8a559a3a6f7 100644 --- a/leanpkg.toml +++ b/leanpkg.toml @@ -1,7 +1,7 @@ [package] name = "mathlib" version = "0.1" -lean_version = "leanprover-community/lean:3.51.0" +lean_version = "leanprover-community/lean:3.51.1" path = "src" [dependencies] From 10878f6bf1dab863445907ab23fbfcefcb5845d0 Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Thu, 25 May 2023 02:59:03 +0000 Subject: [PATCH 1074/1291] feat(linear_algebra/span): generalise to_span_nonzero_singleton (#19082) Co-authored-by: Junyan Xu Co-authored-by: Eric Wieser --- src/linear_algebra/span.lean | 64 ++++++++++++++++-------------------- 1 file changed, 28 insertions(+), 36 deletions(-) diff --git a/src/linear_algebra/span.lean b/src/linear_algebra/span.lean index 54c850ab6fabd..56abfe2cb5545 100644 --- a/src/linear_algebra/span.lean +++ b/src/linear_algebra/span.lean @@ -851,6 +851,15 @@ ext_on hv (set.forall_range_iff.2 h) end add_comm_monoid +section no_zero_divisors + +variables (R M) [ring R] [add_comm_group M] [module R M] [no_zero_smul_divisors R M] + +lemma ker_to_span_singleton {x : M} (h : x ≠ 0) : (to_span_singleton R M x).ker = ⊥ := +set_like.ext $ λ c, smul_eq_zero.trans $ or_iff_left_of_imp $ λ h', (h h').elim + +end no_zero_divisors + section field variables {K V} [field K] [add_comm_group V] [module K V] @@ -867,20 +876,6 @@ eq_top_iff.2 (λ y hy, submodule.mem_sup.2 ⟨(f y * (f x)⁻¹) • x, inv_mul_cancel hx, mul_one, sub_self], by simp only [add_sub_cancel'_right]⟩⟩) -variables (K V) - -lemma ker_to_span_singleton {x : V} (h : x ≠ 0) : (to_span_singleton K V x).ker = ⊥ := -begin - ext c, split, - { intros hc, rw submodule.mem_bot, rw mem_ker at hc, by_contra hc', - have : x = 0, - calc x = c⁻¹ • (c • x) : by rw [← mul_smul, inv_mul_cancel hc', one_smul] - ... = c⁻¹ • ((to_span_singleton K V x) c) : rfl - ... = 0 : by rw [hc, smul_zero], - tauto }, - { rw [mem_ker, submodule.mem_bot], intros h, rw h, simp } -end - end field end linear_map @@ -889,38 +884,35 @@ open linear_map namespace linear_equiv -section field +variables (R M) [ring R] [add_comm_group M] [module R M] [no_zero_smul_divisors R M] + (x : M) (h : x ≠ 0) -variables (K V) [field K] [add_comm_group V] [module K V] - -/-- Given a nonzero element `x` of a vector space `V` over a field `K`, the natural - map from `K` to the span of `x`, with invertibility check to consider it as an - isomorphism.-/ -def to_span_nonzero_singleton (x : V) (h : x ≠ 0) : K ≃ₗ[K] (K ∙ x) := +/-- Given a nonzero element `x` of a torsion-free module `M` over a ring `R`, the natural +isomorphism from `R` to the span of `x` given by $r \mapsto r \cdot x$. -/ +def to_span_nonzero_singleton : R ≃ₗ[R] R ∙ x := linear_equiv.trans (linear_equiv.of_injective - (linear_map.to_span_singleton K V x) (ker_eq_bot.1 $ linear_map.ker_to_span_singleton K V h)) - (linear_equiv.of_eq (to_span_singleton K V x).range (K ∙ x) - (span_singleton_eq_range K V x).symm) + (linear_map.to_span_singleton R M x) (ker_eq_bot.1 $ ker_to_span_singleton R M h)) + (linear_equiv.of_eq (to_span_singleton R M x).range (R ∙ x) + (span_singleton_eq_range R M x).symm) -lemma to_span_nonzero_singleton_one (x : V) (h : x ≠ 0) : - linear_equiv.to_span_nonzero_singleton K V x h 1 = - (⟨x, submodule.mem_span_singleton_self x⟩ : K ∙ x) := +lemma to_span_nonzero_singleton_one : + linear_equiv.to_span_nonzero_singleton R M x h 1 = + (⟨x, submodule.mem_span_singleton_self x⟩ : R ∙ x) := begin apply set_like.coe_eq_coe.mp, - have : ↑(to_span_nonzero_singleton K V x h 1) = to_span_singleton K V x 1 := rfl, + have : ↑(to_span_nonzero_singleton R M x h 1) = to_span_singleton R M x 1 := rfl, rw [this, to_span_singleton_one, submodule.coe_mk], end -/-- Given a nonzero element `x` of a vector space `V` over a field `K`, the natural map - from the span of `x` to `K`.-/ -abbreviation coord (x : V) (h : x ≠ 0) : (K ∙ x) ≃ₗ[K] K := -(to_span_nonzero_singleton K V x h).symm +/-- Given a nonzero element `x` of a torsion-free module `M` over a ring `R`, the natural +isomorphism from the span of `x` to `R` given by $r \cdot x \mapsto r$. -/ +abbreviation coord : (R ∙ x) ≃ₗ[R] R := (to_span_nonzero_singleton R M x h).symm -lemma coord_self (x : V) (h : x ≠ 0) : - (coord K V x h) (⟨x, submodule.mem_span_singleton_self x⟩ : K ∙ x) = 1 := -by rw [← to_span_nonzero_singleton_one K V x h, linear_equiv.symm_apply_apply] +lemma coord_self : (coord R M x h) (⟨x, submodule.mem_span_singleton_self x⟩ : R ∙ x) = 1 := +by rw [← to_span_nonzero_singleton_one R M x h, linear_equiv.symm_apply_apply] -end field +lemma coord_apply_smul (y : submodule.span R ({x} : set M)) : coord R M x h y • x = y := +subtype.ext_iff.1 $ (to_span_nonzero_singleton R M x h).apply_symm_apply _ end linear_equiv From a87d22575d946e1e156fc1edd1e1269600a8a282 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Thu, 25 May 2023 02:59:04 +0000 Subject: [PATCH 1075/1291] chore(*): add mathlib4 synchronization comments (#19089) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.category.Group.filtered_colimits` * `algebra.category.Module.limits` * `algebra.direct_limit` * `algebraic_geometry.prime_spectrum.basic` * `algebraic_geometry.prime_spectrum.is_open_comap_C` * `analysis.calculus.fderiv.star` * `analysis.complex.isometry` * `analysis.locally_convex.with_seminorms` * `analysis.normed.mul_action` * `analysis.normed_space.hahn_banach.extension` * `data.fintype.quotient` * `field_theory.separable_degree` * `linear_algebra.charpoly.to_matrix` * `linear_algebra.matrix.general_linear_group` * `linear_algebra.matrix.to_linear_equiv` * `measure_theory.function.ae_eq_fun` * `model_theory.quotients` * `number_theory.modular_forms.congruence_subgroups` * `ring_theory.adjoin.power_basis` * `ring_theory.adjoin_root` * `ring_theory.laurent_series` * `topology.instances.rat_lemmas` * `topology.metric_space.gromov_hausdorff_realized` --- src/algebra/category/Group/filtered_colimits.lean | 3 +++ src/algebra/category/Module/limits.lean | 3 +++ src/algebra/direct_limit.lean | 3 +++ src/algebraic_geometry/prime_spectrum/basic.lean | 3 +++ src/algebraic_geometry/prime_spectrum/is_open_comap_C.lean | 3 +++ src/analysis/calculus/fderiv/star.lean | 3 +++ src/analysis/complex/isometry.lean | 3 +++ src/analysis/locally_convex/with_seminorms.lean | 3 +++ src/analysis/normed/mul_action.lean | 3 +++ src/analysis/normed_space/hahn_banach/extension.lean | 3 +++ src/data/fintype/quotient.lean | 3 +++ src/field_theory/separable_degree.lean | 3 +++ src/linear_algebra/charpoly/to_matrix.lean | 3 +++ src/linear_algebra/matrix/general_linear_group.lean | 3 +++ src/linear_algebra/matrix/to_linear_equiv.lean | 3 +++ src/measure_theory/function/ae_eq_fun.lean | 3 +++ src/model_theory/quotients.lean | 3 +++ src/number_theory/modular_forms/congruence_subgroups.lean | 3 +++ src/ring_theory/adjoin/power_basis.lean | 3 +++ src/ring_theory/adjoin_root.lean | 3 +++ src/ring_theory/laurent_series.lean | 3 +++ src/topology/instances/rat_lemmas.lean | 3 +++ src/topology/metric_space/gromov_hausdorff_realized.lean | 3 +++ 23 files changed, 69 insertions(+) diff --git a/src/algebra/category/Group/filtered_colimits.lean b/src/algebra/category/Group/filtered_colimits.lean index 7316f5e83bfa6..52be1744411f1 100644 --- a/src/algebra/category/Group/filtered_colimits.lean +++ b/src/algebra/category/Group/filtered_colimits.lean @@ -9,6 +9,9 @@ import algebra.category.Mon.filtered_colimits /-! # The forgetful functor from (commutative) (additive) groups preserves filtered colimits. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Forgetful functors from algebraic categories usually don't preserve colimits. However, they tend to preserve _filtered_ colimits. diff --git a/src/algebra/category/Module/limits.lean b/src/algebra/category/Module/limits.lean index 51386835a918b..43f099089be51 100644 --- a/src/algebra/category/Module/limits.lean +++ b/src/algebra/category/Module/limits.lean @@ -10,6 +10,9 @@ import algebra.direct_limit /-! # The category of R-modules has all limits +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Further, these limits are preserved by the forgetful functor --- that is, the underlying types are just the limits in the category of types. -/ diff --git a/src/algebra/direct_limit.lean b/src/algebra/direct_limit.lean index 891830bcf9987..e30343ee0cee8 100644 --- a/src/algebra/direct_limit.lean +++ b/src/algebra/direct_limit.lean @@ -10,6 +10,9 @@ import ring_theory.ideal.quotient /-! # Direct limit of modules, abelian groups, rings, and fields. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + See Atiyah-Macdonald PP.32-33, Matsumura PP.269-270 Generalizes the notion of "union", or "gluing", of incomparable modules over the same ring, diff --git a/src/algebraic_geometry/prime_spectrum/basic.lean b/src/algebraic_geometry/prime_spectrum/basic.lean index 559f61d56941d..7daf83d95d9ca 100644 --- a/src/algebraic_geometry/prime_spectrum/basic.lean +++ b/src/algebraic_geometry/prime_spectrum/basic.lean @@ -15,6 +15,9 @@ import topology.sober /-! # Prime spectrum of a commutative ring +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The prime spectrum of a commutative ring is the type of all prime ideals. It is naturally endowed with a topology: the Zariski topology. diff --git a/src/algebraic_geometry/prime_spectrum/is_open_comap_C.lean b/src/algebraic_geometry/prime_spectrum/is_open_comap_C.lean index b578d12dbb966..b8802891a5ab3 100644 --- a/src/algebraic_geometry/prime_spectrum/is_open_comap_C.lean +++ b/src/algebraic_geometry/prime_spectrum/is_open_comap_C.lean @@ -6,6 +6,9 @@ Authors: Damiano Testa import algebraic_geometry.prime_spectrum.basic import ring_theory.polynomial.basic /-! +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The morphism `Spec R[x] --> Spec R` induced by the natural inclusion `R --> R[x]` is an open map. The main result is the first part of the statement of Lemma 00FB in the Stacks Project. diff --git a/src/analysis/calculus/fderiv/star.lean b/src/analysis/calculus/fderiv/star.lean index 556f4934f6562..4eaf7d6b21c33 100644 --- a/src/analysis/calculus/fderiv/star.lean +++ b/src/analysis/calculus/fderiv/star.lean @@ -11,6 +11,9 @@ import analysis.normed_space.star.basic /-! # Star operations on derivatives +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + For detailed documentation of the Fréchet derivative, see the module docstring of `analysis/calculus/fderiv/basic.lean`. diff --git a/src/analysis/complex/isometry.lean b/src/analysis/complex/isometry.lean index c20b7eb7f6ec9..44eaf6f25e9c0 100644 --- a/src/analysis/complex/isometry.lean +++ b/src/analysis/complex/isometry.lean @@ -10,6 +10,9 @@ import linear_algebra.matrix.general_linear_group /-! # Isometries of the Complex Plane +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The lemma `linear_isometry_complex` states the classification of isometries in the complex plane. Specifically, isometries with rotations but without translation. The proof involves: diff --git a/src/analysis/locally_convex/with_seminorms.lean b/src/analysis/locally_convex/with_seminorms.lean index a578827117d1d..38e1afa83d665 100644 --- a/src/analysis/locally_convex/with_seminorms.lean +++ b/src/analysis/locally_convex/with_seminorms.lean @@ -12,6 +12,9 @@ import topology.algebra.module.locally_convex /-! # Topology induced by a family of seminorms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions * `seminorm_family.basis_sets`: The set of open seminorm balls for a family of seminorms. diff --git a/src/analysis/normed/mul_action.lean b/src/analysis/normed/mul_action.lean index 5106c76fea567..9765c4e68b77f 100644 --- a/src/analysis/normed/mul_action.lean +++ b/src/analysis/normed/mul_action.lean @@ -9,6 +9,9 @@ import analysis.normed.field.basic /-! # Lemmas for `has_bounded_smul` over normed additive groups +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Lemmas which hold only in `normed_space α β` are provided in another file. Notably we prove that `non_unital_semi_normed_ring`s have bounded actions by left- and right- diff --git a/src/analysis/normed_space/hahn_banach/extension.lean b/src/analysis/normed_space/hahn_banach/extension.lean index f4c581a7f318e..30d4c9802a94c 100644 --- a/src/analysis/normed_space/hahn_banach/extension.lean +++ b/src/analysis/normed_space/hahn_banach/extension.lean @@ -11,6 +11,9 @@ import data.is_R_or_C.lemmas /-! # Extension Hahn-Banach theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove the analytic Hahn-Banach theorem. For any continuous linear function on a subspace, we can extend it to a function on the entire space without changing its norm. diff --git a/src/data/fintype/quotient.lean b/src/data/fintype/quotient.lean index 09fe4cf45be7d..2151d3594a6b3 100644 --- a/src/data/fintype/quotient.lean +++ b/src/data/fintype/quotient.lean @@ -8,6 +8,9 @@ import data.fintype.basic /-! # Quotients of families indexed by a finite type +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file provides `quotient.fin_choice`, a mechanism to go from a finite family of quotients to a quotient of finite families. diff --git a/src/field_theory/separable_degree.lean b/src/field_theory/separable_degree.lean index bc897af345ed3..b0314e1a1c4c0 100644 --- a/src/field_theory/separable_degree.lean +++ b/src/field_theory/separable_degree.lean @@ -11,6 +11,9 @@ import field_theory.separable # Separable degree +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains basics about the separable degree of a polynomial. ## Main results diff --git a/src/linear_algebra/charpoly/to_matrix.lean b/src/linear_algebra/charpoly/to_matrix.lean index b350386871b16..5a4dc9d1a2a3a 100644 --- a/src/linear_algebra/charpoly/to_matrix.lean +++ b/src/linear_algebra/charpoly/to_matrix.lean @@ -11,6 +11,9 @@ import linear_algebra.matrix.basis # Characteristic polynomial +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main result * `linear_map.charpoly_to_matrix f` : `charpoly f` is the characteristic polynomial of the matrix diff --git a/src/linear_algebra/matrix/general_linear_group.lean b/src/linear_algebra/matrix/general_linear_group.lean index 5a6c6f6e35051..659ad251c659e 100644 --- a/src/linear_algebra/matrix/general_linear_group.lean +++ b/src/linear_algebra/matrix/general_linear_group.lean @@ -10,6 +10,9 @@ import linear_algebra.matrix.special_linear_group /-! # The General Linear group $GL(n, R)$ +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the elements of the General Linear group `general_linear_group n R`, consisting of all invertible `n` by `n` `R`-matrices. diff --git a/src/linear_algebra/matrix/to_linear_equiv.lean b/src/linear_algebra/matrix/to_linear_equiv.lean index 18fe1120a77b2..d11f840052bdb 100644 --- a/src/linear_algebra/matrix/to_linear_equiv.lean +++ b/src/linear_algebra/matrix/to_linear_equiv.lean @@ -14,6 +14,9 @@ import ring_theory.localization.integer /-! # Matrices and linear equivalences +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file gives the map `matrix.to_linear_equiv` from matrices with invertible determinant, to linear equivs. diff --git a/src/measure_theory/function/ae_eq_fun.lean b/src/measure_theory/function/ae_eq_fun.lean index a5a408f1d3d7d..85f19720ac1b6 100644 --- a/src/measure_theory/function/ae_eq_fun.lean +++ b/src/measure_theory/function/ae_eq_fun.lean @@ -12,6 +12,9 @@ import measure_theory.function.strongly_measurable.basic # Almost everywhere equal functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We build a space of equivalence classes of functions, where two functions are treated as identical if they are almost everywhere equal. We form the set of equivalence classes under the relation of being almost everywhere equal, which is sometimes known as the `L⁰` space. diff --git a/src/model_theory/quotients.lean b/src/model_theory/quotients.lean index 60ccfe0151a5e..43966d7f0fbe4 100644 --- a/src/model_theory/quotients.lean +++ b/src/model_theory/quotients.lean @@ -8,6 +8,9 @@ import model_theory.semantics /-! # Quotients of First-Order Structures + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. This file defines prestructures and quotients of first-order structures. ## Main Definitions diff --git a/src/number_theory/modular_forms/congruence_subgroups.lean b/src/number_theory/modular_forms/congruence_subgroups.lean index 18d42bbc9f7fa..aa1a2c3e4ba21 100644 --- a/src/number_theory/modular_forms/congruence_subgroups.lean +++ b/src/number_theory/modular_forms/congruence_subgroups.lean @@ -10,6 +10,9 @@ import linear_algebra.matrix.special_linear_group /-! # Congruence subgroups +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This defines congruence subgroups of `SL(2, ℤ)` such as `Γ(N)`, `Γ₀(N)` and `Γ₁(N)` for `N` a natural number. diff --git a/src/ring_theory/adjoin/power_basis.lean b/src/ring_theory/adjoin/power_basis.lean index c4f18ea931fec..070b9fc1506c1 100644 --- a/src/ring_theory/adjoin/power_basis.lean +++ b/src/ring_theory/adjoin/power_basis.lean @@ -11,6 +11,9 @@ import linear_algebra.matrix.basis /-! # Power basis for `algebra.adjoin R {x}` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the canonical power basis on `algebra.adjoin R {x}`, where `x` is an integral element over `R`. -/ diff --git a/src/ring_theory/adjoin_root.lean b/src/ring_theory/adjoin_root.lean index d2c0007771cef..f548e2fc730b2 100644 --- a/src/ring_theory/adjoin_root.lean +++ b/src/ring_theory/adjoin_root.lean @@ -16,6 +16,9 @@ import ring_theory.quotient_noetherian /-! # Adjoining roots of polynomials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the commutative ring `adjoin_root f`, the ring R[X]/(f) obtained from a commutative ring `R` and a polynomial `f : R[X]`. If furthermore `R` is a field and `f` is irreducible, the field structure on `adjoin_root f` is constructed. diff --git a/src/ring_theory/laurent_series.lean b/src/ring_theory/laurent_series.lean index 137a2f03ff5a7..430e199e496ea 100644 --- a/src/ring_theory/laurent_series.lean +++ b/src/ring_theory/laurent_series.lean @@ -9,6 +9,9 @@ import ring_theory.localization.fraction_ring /-! # Laurent Series +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main Definitions * Defines `laurent_series` as an abbreviation for `hahn_series ℤ`. * Provides a coercion `power_series R` into `laurent_series R` given by diff --git a/src/topology/instances/rat_lemmas.lean b/src/topology/instances/rat_lemmas.lean index fe0059dec2da3..b20d5f7e8d5a2 100644 --- a/src/topology/instances/rat_lemmas.lean +++ b/src/topology/instances/rat_lemmas.lean @@ -10,6 +10,9 @@ import topology.alexandroff /-! # Additional lemmas about the topology on rational numbers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The structure of a metric space on `ℚ` (`rat.metric_space`) is introduced elsewhere, induced from `ℝ`. In this file we prove some properties of this topological space and its one-point compactification. diff --git a/src/topology/metric_space/gromov_hausdorff_realized.lean b/src/topology/metric_space/gromov_hausdorff_realized.lean index 0a25b1becc138..d4c8ea3a88b51 100644 --- a/src/topology/metric_space/gromov_hausdorff_realized.lean +++ b/src/topology/metric_space/gromov_hausdorff_realized.lean @@ -10,6 +10,9 @@ import topology.continuous_function.bounded /-! # The Gromov-Hausdorff distance is realized +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we construct of a good coupling between nonempty compact metric spaces, minimizing their Hausdorff distance. This construction is instrumental to study the Gromov-Hausdorff distance between nonempty compact metric spaces. From 6480bedf0354aaba0016eeefd726f7e3e5fc50aa Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 25 May 2023 17:14:54 +0000 Subject: [PATCH 1076/1291] chore(measure_theory/function/lp_space): add nnnorm lemmas (#19091) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This also adds a `has_nnnorm (Lp E p μ)` instance, which applies more generally (without `[fact (1 ≤ p)]`) than the version derived from `Lp.normed_add_comm_group`. In general, `nnnorm` (`‖f x‖₊)` statements can often be easier to work with because there is no need to repeatedly remind Lean that the norm is non-negative. Most of the time, the `norm` (`‖x‖`) counterparts follow trivially from the `nnnorm` ones. Notably this removes the need for a proof-by-cases in `snorm_le_mul_snorm_of_ae_le_mul` and some similar lemmas. --- src/measure_theory/function/lp_space.lean | 256 +++++++++++++++------- 1 file changed, 179 insertions(+), 77 deletions(-) diff --git a/src/measure_theory/function/lp_space.lean b/src/measure_theory/function/lp_space.lean index 8ce0cf99e6b55..918ddb4fa3478 100644 --- a/src/measure_theory/function/lp_space.lean +++ b/src/measure_theory/function/lp_space.lean @@ -359,7 +359,7 @@ end end const -lemma snorm'_mono_ae {f : α → F} {g : α → G} (hq : 0 ≤ q) (h : ∀ᵐ x ∂μ, ‖f x‖ ≤ ‖g x‖) : +lemma snorm'_mono_nnnorm_ae {f : α → F} {g : α → G} (hq : 0 ≤ q) (h : ∀ᵐ x ∂μ, ‖f x‖₊ ≤ ‖g x‖₊) : snorm' f q μ ≤ snorm' g q μ := begin rw [snorm'], @@ -368,75 +368,103 @@ begin exact ennreal.rpow_le_rpow (ennreal.coe_le_coe.2 hx) hq end -lemma snorm'_congr_norm_ae {f g : α → F} (hfg : ∀ᵐ x ∂μ, ‖f x‖ = ‖g x‖) : +lemma snorm'_mono_ae {f : α → F} {g : α → G} (hq : 0 ≤ q) (h : ∀ᵐ x ∂μ, ‖f x‖ ≤ ‖g x‖) : + snorm' f q μ ≤ snorm' g q μ := +snorm'_mono_nnnorm_ae hq h + +lemma snorm'_congr_nnnorm_ae {f g : α → F} (hfg : ∀ᵐ x ∂μ, ‖f x‖₊ = ‖g x‖₊) : snorm' f q μ = snorm' g q μ := begin have : (λ x, (‖f x‖₊ ^ q : ℝ≥0∞)) =ᵐ[μ] (λ x, ‖g x‖₊ ^ q), - from hfg.mono (λ x hx, by { simp only [← coe_nnnorm, nnreal.coe_eq] at hx, simp [hx] }), + from hfg.mono (λ x hx, by simp_rw hx), simp only [snorm', lintegral_congr_ae this] end +lemma snorm'_congr_norm_ae {f g : α → F} (hfg : ∀ᵐ x ∂μ, ‖f x‖ = ‖g x‖) : + snorm' f q μ = snorm' g q μ := +snorm'_congr_nnnorm_ae $ hfg.mono $ λ x hx, nnreal.eq hx + lemma snorm'_congr_ae {f g : α → F} (hfg : f =ᵐ[μ] g) : snorm' f q μ = snorm' g q μ := -snorm'_congr_norm_ae (hfg.fun_comp _) +snorm'_congr_nnnorm_ae (hfg.fun_comp _) lemma snorm_ess_sup_congr_ae {f g : α → F} (hfg : f =ᵐ[μ] g) : snorm_ess_sup f μ = snorm_ess_sup g μ := ess_sup_congr_ae (hfg.fun_comp (coe ∘ nnnorm)) -lemma snorm_mono_ae {f : α → F} {g : α → G} (h : ∀ᵐ x ∂μ, ‖f x‖ ≤ ‖g x‖) : +lemma snorm_mono_nnnorm_ae {f : α → F} {g : α → G} (h : ∀ᵐ x ∂μ, ‖f x‖₊ ≤ ‖g x‖₊) : snorm f p μ ≤ snorm g p μ := begin simp only [snorm], split_ifs, { exact le_rfl }, - { refine ess_sup_mono_ae (h.mono $ λ x hx, _), - exact_mod_cast hx }, - { exact snorm'_mono_ae ennreal.to_real_nonneg h } + { exact ess_sup_mono_ae (h.mono $ λ x hx, ennreal.coe_le_coe.mpr hx) }, + { exact snorm'_mono_nnnorm_ae ennreal.to_real_nonneg h } end +lemma snorm_mono_ae {f : α → F} {g : α → G} (h : ∀ᵐ x ∂μ, ‖f x‖ ≤ ‖g x‖) : + snorm f p μ ≤ snorm g p μ := +snorm_mono_nnnorm_ae h + lemma snorm_mono_ae_real {f : α → F} {g : α → ℝ} (h : ∀ᵐ x ∂μ, ‖f x‖ ≤ g x) : snorm f p μ ≤ snorm g p μ := snorm_mono_ae $ h.mono (λ x hx, hx.trans ((le_abs_self _).trans (real.norm_eq_abs _).symm.le)) +lemma snorm_mono_nnnorm {f : α → F} {g : α → G} (h : ∀ x, ‖f x‖₊ ≤ ‖g x‖₊) : + snorm f p μ ≤ snorm g p μ := +snorm_mono_nnnorm_ae (eventually_of_forall (λ x, h x)) + lemma snorm_mono {f : α → F} {g : α → G} (h : ∀ x, ‖f x‖ ≤ ‖g x‖) : snorm f p μ ≤ snorm g p μ := -snorm_mono_ae (eventually_of_forall (λ x, h x)) +snorm_mono_nnnorm h lemma snorm_mono_real {f : α → F} {g : α → ℝ} (h : ∀ x, ‖f x‖ ≤ g x) : snorm f p μ ≤ snorm g p μ := snorm_mono_ae_real (eventually_of_forall (λ x, h x)) +lemma snorm_ess_sup_le_of_ae_nnnorm_bound {f : α → F} {C : ℝ≥0} (hfC : ∀ᵐ x ∂μ, ‖f x‖₊ ≤ C) : + snorm_ess_sup f μ ≤ C := +ess_sup_le_of_ae_le C $ hfC.mono $ λ x hx, ennreal.coe_le_coe.mpr hx + lemma snorm_ess_sup_le_of_ae_bound {f : α → F} {C : ℝ} (hfC : ∀ᵐ x ∂μ, ‖f x‖ ≤ C) : - snorm_ess_sup f μ ≤ ennreal.of_real C:= -begin - simp_rw [snorm_ess_sup, ← of_real_norm_eq_coe_nnnorm], - refine ess_sup_le_of_ae_le (ennreal.of_real C) (hfC.mono (λ x hx, _)), - exact ennreal.of_real_le_of_real hx, -end + snorm_ess_sup f μ ≤ ennreal.of_real C := +snorm_ess_sup_le_of_ae_nnnorm_bound $ hfC.mono $ λ x hx, hx.trans C.le_coe_to_nnreal + +lemma snorm_ess_sup_lt_top_of_ae_nnnorm_bound {f : α → F} {C : ℝ≥0} (hfC : ∀ᵐ x ∂μ, ‖f x‖₊ ≤ C) : + snorm_ess_sup f μ < ∞ := +(snorm_ess_sup_le_of_ae_nnnorm_bound hfC).trans_lt ennreal.coe_lt_top lemma snorm_ess_sup_lt_top_of_ae_bound {f : α → F} {C : ℝ} (hfC : ∀ᵐ x ∂μ, ‖f x‖ ≤ C) : snorm_ess_sup f μ < ∞ := (snorm_ess_sup_le_of_ae_bound hfC).trans_lt ennreal.of_real_lt_top -lemma snorm_le_of_ae_bound {f : α → F} {C : ℝ} (hfC : ∀ᵐ x ∂μ, ‖f x‖ ≤ C) : - snorm f p μ ≤ ((μ set.univ) ^ p.to_real⁻¹) * (ennreal.of_real C) := +lemma snorm_le_of_ae_nnnorm_bound {f : α → F} {C : ℝ≥0} (hfC : ∀ᵐ x ∂μ, ‖f x‖₊ ≤ C) : + snorm f p μ ≤ C • (μ set.univ ^ p.to_real⁻¹) := begin by_cases hμ : μ = 0, { simp [hμ] }, haveI : μ.ae.ne_bot := ae_ne_bot.mpr hμ, by_cases hp : p = 0, { simp [hp] }, - have hC : 0 ≤ C, from le_trans (norm_nonneg _) hfC.exists.some_spec, - have hC' : ‖C‖ = C := by rw [real.norm_eq_abs, abs_eq_self.mpr hC], - have : ∀ᵐ x ∂μ, ‖f x‖ ≤ ‖(λ _, C) x‖, from hfC.mono (λ x hx, hx.trans (le_of_eq hC'.symm)), - convert snorm_mono_ae this, - rw [snorm_const _ hp hμ, mul_comm, ← of_real_norm_eq_coe_nnnorm, hC', one_div] + have : ∀ᵐ x ∂μ, ‖f x‖₊ ≤ ‖(C : ℝ)‖₊ := hfC.mono (λ x hx, hx.trans_eq C.nnnorm_eq.symm), + refine (snorm_mono_ae this).trans_eq _, + rw [snorm_const _ hp hμ, C.nnnorm_eq, one_div, ennreal.smul_def, smul_eq_mul], end +lemma snorm_le_of_ae_bound {f : α → F} {C : ℝ} (hfC : ∀ᵐ x ∂μ, ‖f x‖ ≤ C) : + snorm f p μ ≤ ((μ set.univ) ^ p.to_real⁻¹) * (ennreal.of_real C) := +begin + rw [←mul_comm], + exact snorm_le_of_ae_nnnorm_bound (hfC.mono $ λ x hx, hx.trans C.le_coe_to_nnreal), +end + +lemma snorm_congr_nnnorm_ae {f : α → F} {g : α → G} (hfg : ∀ᵐ x ∂μ, ‖f x‖₊ = ‖g x‖₊) : + snorm f p μ = snorm g p μ := +le_antisymm (snorm_mono_nnnorm_ae $ eventually_eq.le hfg) + (snorm_mono_nnnorm_ae $ (eventually_eq.symm hfg).le) + lemma snorm_congr_norm_ae {f : α → F} {g : α → G} (hfg : ∀ᵐ x ∂μ, ‖f x‖ = ‖g x‖) : snorm f p μ = snorm g p μ := -le_antisymm (snorm_mono_ae $ eventually_eq.le hfg) - (snorm_mono_ae $ (eventually_eq.symm hfg).le) +snorm_congr_nnnorm_ae $ hfg.mono $ λ x hx, nnreal.eq hx @[simp] lemma snorm'_norm {f : α → F} : snorm' (λ a, ‖f a‖) q μ = snorm' f q μ := by simp [snorm'] @@ -1389,43 +1417,56 @@ end normed_space section monotonicity -lemma snorm_le_mul_snorm_aux_of_nonneg {f : α → F} {g : α → G} {c : ℝ} - (h : ∀ᵐ x ∂μ, ‖f x‖ ≤ c * ‖g x‖) (hc : 0 ≤ c) (p : ℝ≥0∞) : - snorm f p μ ≤ (ennreal.of_real c) * snorm g p μ := +lemma snorm_le_nnreal_smul_snorm_of_ae_le_mul {f : α → F} {g : α → G} {c : ℝ≥0} + (h : ∀ᵐ x ∂μ, ‖f x‖₊ ≤ c * ‖g x‖₊) (p : ℝ≥0∞) : + snorm f p μ ≤ c • snorm g p μ := begin - lift c to ℝ≥0 using hc, - rw [ennreal.of_real_coe_nnreal, ← c.nnnorm_eq, ← snorm_norm g, ← snorm_const_smul (c : ℝ)], + rw [← snorm_norm g, ←c.nnnorm_eq, ennreal.smul_def, smul_eq_mul, ← snorm_const_smul (c : ℝ)], swap, apply_instance, - refine snorm_mono_ae _, - simpa + refine snorm_mono_nnnorm_ae (h.mono $ λ x hx, hx.trans_eq _), + rw [pi.smul_apply, smul_eq_mul, nnnorm_mul, nnreal.nnnorm_eq, nnnorm_norm], +end + +-- TODO: add the whole family of lemmas? +private lemma le_mul_iff_eq_zero_of_nonneg_of_neg_of_nonneg {α} [linear_ordered_semiring α] + {a b c : α} (ha : 0 ≤ a) (hb : b < 0) (hc : 0 ≤ c) : a ≤ b * c ↔ a = 0 ∧ c = 0 := +begin + split, + { intro h, + exact ⟨(h.trans (mul_nonpos_of_nonpos_of_nonneg hb.le hc)).antisymm ha, + (nonpos_of_mul_nonneg_right (ha.trans h) hb).antisymm hc⟩ }, + { rintro ⟨rfl, rfl⟩, + rw mul_zero, } end -lemma snorm_le_mul_snorm_aux_of_neg {f : α → F} {g : α → G} {c : ℝ} +/-- When `c` is negative, `‖f x‖ ≤ c * ‖g x‖` is nonsense and forces both `f` and `g` to have an +`snorm` of `0`. -/ +lemma snorm_eq_zero_and_zero_of_ae_le_mul_neg {f : α → F} {g : α → G} {c : ℝ} (h : ∀ᵐ x ∂μ, ‖f x‖ ≤ c * ‖g x‖) (hc : c < 0) (p : ℝ≥0∞) : snorm f p μ = 0 ∧ snorm g p μ = 0 := begin - suffices : f =ᵐ[μ] 0 ∧ g =ᵐ[μ] 0, - by simp [snorm_congr_ae this.1, snorm_congr_ae this.2], - refine ⟨h.mono $ λ x hx, _, h.mono $ λ x hx, _⟩, - { refine norm_le_zero_iff.1 (hx.trans _), - exact mul_nonpos_of_nonpos_of_nonneg hc.le (norm_nonneg _) }, - { refine norm_le_zero_iff.1 (nonpos_of_mul_nonneg_right _ hc), - exact (norm_nonneg _).trans hx } + simp_rw [le_mul_iff_eq_zero_of_nonneg_of_neg_of_nonneg (norm_nonneg _) hc (norm_nonneg _), + norm_eq_zero, eventually_and] at h, + change f =ᵐ[μ] 0 ∧ g =ᵐ[μ] 0 at h, + simp [snorm_congr_ae h.1, snorm_congr_ae h.2], end lemma snorm_le_mul_snorm_of_ae_le_mul {f : α → F} {g : α → G} {c : ℝ} (h : ∀ᵐ x ∂μ, ‖f x‖ ≤ c * ‖g x‖) (p : ℝ≥0∞) : snorm f p μ ≤ (ennreal.of_real c) * snorm g p μ := -begin - cases le_or_lt 0 c with hc hc, - { exact snorm_le_mul_snorm_aux_of_nonneg h hc p }, - { simp [snorm_le_mul_snorm_aux_of_neg h hc p] } -end +snorm_le_nnreal_smul_snorm_of_ae_le_mul + (h.mono $ λ x hx, hx.trans $ mul_le_mul_of_nonneg_right c.le_coe_to_nnreal (norm_nonneg _)) _ + +lemma mem_ℒp.of_nnnorm_le_mul {f : α → E} {g : α → F} {c : ℝ≥0} + (hg : mem_ℒp g p μ) (hf : ae_strongly_measurable f μ) (hfg : ∀ᵐ x ∂μ, ‖f x‖₊ ≤ c * ‖g x‖₊) : + mem_ℒp f p μ := +⟨hf, (snorm_le_nnreal_smul_snorm_of_ae_le_mul hfg p).trans_lt $ + ennreal.mul_lt_top ennreal.coe_ne_top hg.snorm_ne_top⟩ lemma mem_ℒp.of_le_mul {f : α → E} {g : α → F} {c : ℝ} (hg : mem_ℒp g p μ) (hf : ae_strongly_measurable f μ) (hfg : ∀ᵐ x ∂μ, ‖f x‖ ≤ c * ‖g x‖) : mem_ℒp f p μ := -⟨hf, lt_of_le_of_lt (snorm_le_mul_snorm_of_ae_le_mul hfg p) $ +⟨hf, (snorm_le_mul_snorm_of_ae_le_mul hfg p).trans_lt $ ennreal.mul_lt_top ennreal.of_real_ne_top hg.snorm_ne_top⟩ end monotonicity @@ -1651,16 +1692,28 @@ lemma mem_Lp_const (α) {m : measurable_space α} (μ : measure α) (c : E) [is_ instance : has_norm (Lp E p μ) := { norm := λ f, ennreal.to_real (snorm f p μ) } +-- note: we need this to be defeq to the instance from `seminormed_add_group.to_has_nnnorm`, so +-- can't use `ennreal.to_nnreal (snorm f p μ)` +instance : has_nnnorm (Lp E p μ) := { nnnorm := λ f, ⟨‖f‖, ennreal.to_real_nonneg⟩ } + instance : has_dist (Lp E p μ) := { dist := λ f g, ‖f - g‖} instance : has_edist (Lp E p μ) := { edist := λ f g, snorm (f - g) p μ } lemma norm_def (f : Lp E p μ) : ‖f‖ = ennreal.to_real (snorm f p μ) := rfl +lemma nnnorm_def (f : Lp E p μ) : ‖f‖₊ = ennreal.to_nnreal (snorm f p μ) := subtype.eta _ _ + +@[simp, norm_cast] protected lemma coe_nnnorm (f : Lp E p μ) : (‖f‖₊ : ℝ) = ‖f‖ := rfl + @[simp] lemma norm_to_Lp (f : α → E) (hf : mem_ℒp f p μ) : ‖hf.to_Lp f‖ = ennreal.to_real (snorm f p μ) := by rw [norm_def, snorm_congr_ae (mem_ℒp.coe_fn_to_Lp hf)] +@[simp] lemma nnnorm_to_Lp (f : α → E) (hf : mem_ℒp f p μ) : + ‖hf.to_Lp f‖₊ = ennreal.to_nnreal (snorm f p μ) := +nnreal.eq $ norm_to_Lp f hf + lemma dist_def (f g : Lp E p μ) : dist f g = (snorm (f - g) p μ).to_real := begin simp_rw [dist, norm_def], @@ -1679,22 +1732,29 @@ by { rw edist_def, exact snorm_congr_ae (hf.coe_fn_to_Lp.sub hg.coe_fn_to_Lp) } edist (hf.to_Lp f) 0 = snorm f p μ := by { convert edist_to_Lp_to_Lp f 0 hf zero_mem_ℒp, simp } -@[simp] lemma norm_zero : ‖(0 : Lp E p μ)‖ = 0 := +@[simp] lemma nnnorm_zero : ‖(0 : Lp E p μ)‖₊ = 0 := begin - change (snorm ⇑(0 : α →ₘ[μ] E) p μ).to_real = 0, + rw [nnnorm_def], + change (snorm ⇑(0 : α →ₘ[μ] E) p μ).to_nnreal = 0, simp [snorm_congr_ae ae_eq_fun.coe_fn_zero, snorm_zero] end -lemma norm_eq_zero_iff {f : Lp E p μ} (hp : 0 < p) : ‖f‖ = 0 ↔ f = 0 := +@[simp] lemma norm_zero : ‖(0 : Lp E p μ)‖ = 0 := +congr_arg coe nnnorm_zero + +lemma nnnorm_eq_zero_iff {f : Lp E p μ} (hp : 0 < p) : ‖f‖₊ = 0 ↔ f = 0 := begin refine ⟨λ hf, _, λ hf, by simp [hf]⟩, - rw [norm_def, ennreal.to_real_eq_zero_iff] at hf, + rw [nnnorm_def, ennreal.to_nnreal_eq_zero_iff] at hf, cases hf, { rw snorm_eq_zero_iff (Lp.ae_strongly_measurable f) hp.ne.symm at hf, exact subtype.eq (ae_eq_fun.ext (hf.trans ae_eq_fun.coe_fn_zero.symm)), }, { exact absurd hf (snorm_ne_top f), }, end +lemma norm_eq_zero_iff {f : Lp E p μ} (hp : 0 < p) : ‖f‖ = 0 ↔ f = 0 := +iff.symm $ (nnnorm_eq_zero_iff hp).symm.trans $ (nnreal.coe_eq_zero _).symm + lemma eq_zero_iff_ae_eq_zero {f : Lp E p μ} : f = 0 ↔ f =ᵐ[μ] 0 := begin split, @@ -1708,19 +1768,31 @@ begin exact h'a.symm, }, end -@[simp] lemma norm_neg {f : Lp E p μ} : ‖-f‖ = ‖f‖ := -by rw [norm_def, norm_def, snorm_congr_ae (coe_fn_neg _), snorm_neg] +@[simp] lemma nnnorm_neg (f : Lp E p μ) : ‖-f‖₊ = ‖f‖₊ := +by rw [nnnorm_def, nnnorm_def, snorm_congr_ae (coe_fn_neg _), snorm_neg] + +@[simp] lemma norm_neg (f : Lp E p μ) : ‖-f‖ = ‖f‖ := +(congr_arg (coe : ℝ≥0 → ℝ) (nnnorm_neg f) : _) + +lemma nnnorm_le_mul_nnnorm_of_ae_le_mul {c : ℝ≥0} {f : Lp E p μ} {g : Lp F p μ} + (h : ∀ᵐ x ∂μ, ‖f x‖₊ ≤ c * ‖g x‖₊ ) : ‖f‖₊ ≤ c * ‖g‖₊ := +begin + simp only [nnnorm_def], + have := snorm_le_nnreal_smul_snorm_of_ae_le_mul h p, + rwa [← ennreal.to_nnreal_le_to_nnreal, ennreal.smul_def, smul_eq_mul, ennreal.to_nnreal_mul, + ennreal.to_nnreal_coe] at this, + { exact (Lp.mem_ℒp _).snorm_ne_top }, + { exact ennreal.mul_ne_top ennreal.coe_ne_top (Lp.mem_ℒp _).snorm_ne_top }, +end lemma norm_le_mul_norm_of_ae_le_mul {c : ℝ} {f : Lp E p μ} {g : Lp F p μ} (h : ∀ᵐ x ∂μ, ‖f x‖ ≤ c * ‖g x‖) : ‖f‖ ≤ c * ‖g‖ := begin - simp only [norm_def], cases le_or_lt 0 c with hc hc, - { have := snorm_le_mul_snorm_of_ae_le_mul h p, - rwa [← ennreal.to_real_le_to_real, ennreal.to_real_mul, ennreal.to_real_of_real hc] at this, - { exact (Lp.mem_ℒp _).snorm_ne_top }, - { exact ennreal.mul_ne_top ennreal.of_real_ne_top (Lp.mem_ℒp _).snorm_ne_top } }, - { have := snorm_le_mul_snorm_aux_of_neg h hc p, + { lift c to ℝ≥0 using hc, + exact nnreal.coe_le_coe.mpr (nnnorm_le_mul_nnnorm_of_ae_le_mul h) }, + { simp only [norm_def], + have := snorm_eq_zero_and_zero_of_ae_le_mul_neg h hc p, simp [this] } end @@ -1731,34 +1803,53 @@ begin exact snorm_mono_ae h end +lemma mem_Lp_of_nnnorm_ae_le_mul {c : ℝ≥0} {f : α →ₘ[μ] E} {g : Lp F p μ} + (h : ∀ᵐ x ∂μ, ‖f x‖₊ ≤ c * ‖g x‖₊) : + f ∈ Lp E p μ := +mem_Lp_iff_mem_ℒp.2 $ mem_ℒp.of_nnnorm_le_mul (Lp.mem_ℒp g) f.ae_strongly_measurable h + lemma mem_Lp_of_ae_le_mul {c : ℝ} {f : α →ₘ[μ] E} {g : Lp F p μ} (h : ∀ᵐ x ∂μ, ‖f x‖ ≤ c * ‖g x‖) : f ∈ Lp E p μ := mem_Lp_iff_mem_ℒp.2 $ mem_ℒp.of_le_mul (Lp.mem_ℒp g) f.ae_strongly_measurable h -lemma mem_Lp_of_ae_le {f : α →ₘ[μ] E} {g : Lp F p μ} (h : ∀ᵐ x ∂μ, ‖f x‖ ≤ ‖g x‖) : +lemma mem_Lp_of_nnnorm_ae_le {f : α →ₘ[μ] E} {g : Lp F p μ} (h : ∀ᵐ x ∂μ, ‖f x‖₊ ≤ ‖g x‖₊) : f ∈ Lp E p μ := mem_Lp_iff_mem_ℒp.2 $ mem_ℒp.of_le (Lp.mem_ℒp g) f.ae_strongly_measurable h +lemma mem_Lp_of_ae_le {f : α →ₘ[μ] E} {g : Lp F p μ} (h : ∀ᵐ x ∂μ, ‖f x‖ ≤ ‖g x‖) : + f ∈ Lp E p μ := +mem_Lp_of_nnnorm_ae_le h + +lemma mem_Lp_of_ae_nnnorm_bound [is_finite_measure μ] {f : α →ₘ[μ] E} (C : ℝ≥0) + (hfC : ∀ᵐ x ∂μ, ‖f x‖₊ ≤ C) : + f ∈ Lp E p μ := +mem_Lp_iff_mem_ℒp.2 $ mem_ℒp.of_bound f.ae_strongly_measurable _ hfC + lemma mem_Lp_of_ae_bound [is_finite_measure μ] {f : α →ₘ[μ] E} (C : ℝ) (hfC : ∀ᵐ x ∂μ, ‖f x‖ ≤ C) : f ∈ Lp E p μ := mem_Lp_iff_mem_ℒp.2 $ mem_ℒp.of_bound f.ae_strongly_measurable _ hfC -lemma norm_le_of_ae_bound [is_finite_measure μ] {f : Lp E p μ} {C : ℝ} (hC : 0 ≤ C) - (hfC : ∀ᵐ x ∂μ, ‖f x‖ ≤ C) : - ‖f‖ ≤ (measure_univ_nnreal μ) ^ (p.to_real)⁻¹ * C := +lemma nnnorm_le_of_ae_bound [is_finite_measure μ] {f : Lp E p μ} {C : ℝ≥0} + (hfC : ∀ᵐ x ∂μ, ‖f x‖₊ ≤ C) : + ‖f‖₊ ≤ (measure_univ_nnreal μ) ^ (p.to_real)⁻¹ * C := begin by_cases hμ : μ = 0, { by_cases hp : p.to_real⁻¹ = 0, - { simpa [hp, hμ, norm_def] using hC }, - { simp [hμ, norm_def, real.zero_rpow hp] } }, - let A : ℝ≥0 := (measure_univ_nnreal μ) ^ (p.to_real)⁻¹ * ⟨C, hC⟩, - suffices : snorm f p μ ≤ A, - { exact ennreal.to_real_le_coe_of_le_coe this }, - convert snorm_le_of_ae_bound hfC, + { simp [hp, hμ, nnnorm_def] }, + { simp [hμ, nnnorm_def, real.zero_rpow hp] } }, + rw [←ennreal.coe_le_coe, nnnorm_def, ennreal.coe_to_nnreal (snorm_ne_top _)], + refine (snorm_le_of_ae_nnnorm_bound hfC).trans_eq _, rw [← coe_measure_univ_nnreal μ, ennreal.coe_rpow_of_ne_zero (measure_univ_nnreal_pos hμ).ne', - ennreal.coe_mul], - congr, - rw max_eq_left hC + ennreal.coe_mul, mul_comm, ennreal.smul_def, smul_eq_mul], +end + +lemma norm_le_of_ae_bound [is_finite_measure μ] {f : Lp E p μ} {C : ℝ} (hC : 0 ≤ C) + (hfC : ∀ᵐ x ∂μ, ‖f x‖ ≤ C) : + ‖f‖ ≤ (measure_univ_nnreal μ) ^ (p.to_real)⁻¹ * C := +begin + lift C to ℝ≥0 using hC, + have := nnnorm_le_of_ae_bound hfC, + rwa [←nnreal.coe_le_coe, nnreal.coe_mul, nnreal.coe_rpow] at this, end instance [hp : fact (1 ≤ p)] : normed_add_comm_group (Lp E p μ) := @@ -1786,6 +1877,10 @@ example [fact (1 ≤ p)] : pseudo_emetric_space.to_has_edist = (Lp.has_edist : has_edist (Lp E p μ)) := rfl +example [fact (1 ≤ p)] : + seminormed_add_group.to_has_nnnorm = (Lp.has_nnnorm : has_nnnorm (Lp E p μ)) := +rfl + section normed_space variables {𝕜 : Type*} [normed_field 𝕜] [normed_space 𝕜 E] @@ -2877,18 +2972,25 @@ begin convert f.norm_coe_le_norm x end +/-- The `Lp`-norm of a bounded continuous function is at most a constant (depending on the measure +of the whole space) times its sup-norm. -/ +lemma Lp_nnnorm_le (f : α →ᵇ E) : + ‖(⟨f.to_continuous_map.to_ae_eq_fun μ, mem_Lp f⟩ : Lp E p μ)‖₊ + ≤ (measure_univ_nnreal μ) ^ (p.to_real)⁻¹ * ‖f‖₊ := +begin + apply Lp.nnnorm_le_of_ae_bound, + refine (f.to_continuous_map.coe_fn_to_ae_eq_fun μ).mono _, + intros x hx, + rw [←nnreal.coe_le_coe, coe_nnnorm, coe_nnnorm], + convert f.norm_coe_le_norm x, +end + /-- The `Lp`-norm of a bounded continuous function is at most a constant (depending on the measure of the whole space) times its sup-norm. -/ lemma Lp_norm_le (f : α →ᵇ E) : ‖(⟨f.to_continuous_map.to_ae_eq_fun μ, mem_Lp f⟩ : Lp E p μ)‖ ≤ (measure_univ_nnreal μ) ^ (p.to_real)⁻¹ * ‖f‖ := -begin - apply Lp.norm_le_of_ae_bound (norm_nonneg f), - { refine (f.to_continuous_map.coe_fn_to_ae_eq_fun μ).mono _, - intros x hx, - convert f.norm_coe_le_norm x }, - { apply_instance } -end +Lp_nnnorm_le f variables (p μ) From de83b43717abe353f425855fcf0cedf9ea0fe8a4 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 25 May 2023 18:32:44 +0000 Subject: [PATCH 1077/1291] feat(analysis/normed_space/lp_space): generalize from normed_field to `normed_ring` (#19085) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This provides a `module 𝕜 (lp E p)` instance for `normed_ring 𝕜` instead of `normed_field 𝕜`, as we didn't actually require that the bound on the norm of the scalar action was tight. --- src/analysis/normed_space/lp_space.lean | 132 ++++++++++++++++-------- 1 file changed, 91 insertions(+), 41 deletions(-) diff --git a/src/analysis/normed_space/lp_space.lean b/src/analysis/normed_space/lp_space.lean index 7ff868c4f74af..4d69acab9ca0d 100644 --- a/src/analysis/normed_space/lp_space.lean +++ b/src/analysis/normed_space/lp_space.lean @@ -248,9 +248,9 @@ begin exact (hf i (s.mem_insert_self i)).add (ih (λ j hj, hf j (finset.mem_insert_of_mem hj))), }, end -section normed_space +section has_bounded_smul -variables {𝕜 : Type*} [normed_field 𝕜] [Π i, normed_space 𝕜 (E i)] +variables {𝕜 : Type*} [normed_ring 𝕜] [Π i, module 𝕜 (E i)] [∀ i, has_bounded_smul 𝕜 (E i)] lemma const_smul {f : Π i, E i} (hf : mem_ℓp f p) (c : 𝕜) : mem_ℓp (c • f) p := begin @@ -261,17 +261,21 @@ begin { obtain ⟨A, hA⟩ := hf.bdd_above, refine mem_ℓp_infty ⟨‖c‖ * A, _⟩, rintros a ⟨i, rfl⟩, - simpa [norm_smul] using mul_le_mul_of_nonneg_left (hA ⟨i, rfl⟩) (norm_nonneg c) }, + refine (norm_smul_le _ _).trans _, + exact mul_le_mul_of_nonneg_left (hA ⟨i, rfl⟩) (norm_nonneg c) }, { apply mem_ℓp_gen, - convert (hf.summable hp).mul_left (‖c‖ ^ p.to_real), - ext i, - simp [norm_smul, real.mul_rpow (norm_nonneg c) (norm_nonneg (f i))] }, + have := (hf.summable hp).mul_left (↑(‖c‖₊ ^ p.to_real) : ℝ), + simp_rw [← coe_nnnorm, ←nnreal.coe_rpow, ←nnreal.coe_mul, nnreal.summable_coe, + ←nnreal.mul_rpow] at this ⊢, + refine nnreal.summable_of_le _ this, + intro i, + exact nnreal.rpow_le_rpow (nnnorm_smul_le _ _) (ennreal.to_real_nonneg), }, end lemma const_mul {f : α → 𝕜} (hf : mem_ℓp f p) (c : 𝕜) : mem_ℓp (λ x, c * f x) p := -@mem_ℓp.const_smul α (λ i, 𝕜) _ _ 𝕜 _ _ _ hf c +@mem_ℓp.const_smul α (λ i, 𝕜) _ _ 𝕜 _ _ (λ i, by apply_instance) _ hf c -end normed_space +end has_bounded_smul end mem_ℓp @@ -562,12 +566,25 @@ norm_le_of_tsum_le hp hC (tsum_le_of_sum_le ((lp.mem_ℓp f).summable hp) hf) end compare_pointwise -section normed_space - -variables {𝕜 : Type*} [normed_field 𝕜] [Π i, normed_space 𝕜 (E i)] +section has_bounded_smul +variables {𝕜 : Type*} {𝕜' : Type*} +variables [normed_ring 𝕜] [normed_ring 𝕜'] +variables [Π i, module 𝕜 (E i)] [Π i, module 𝕜' (E i)] instance : module 𝕜 (pre_lp E) := pi.module α E 𝕜 +instance [Π i, smul_comm_class 𝕜' 𝕜 (E i)] : smul_comm_class 𝕜' 𝕜 (pre_lp E) := +pi.smul_comm_class + +instance [has_smul 𝕜' 𝕜] [Π i, is_scalar_tower 𝕜' 𝕜 (E i)] : is_scalar_tower 𝕜' 𝕜 (pre_lp E) := +pi.is_scalar_tower + +instance [Π i, module 𝕜ᵐᵒᵖ (E i)] [Π i, is_central_scalar 𝕜 (E i)] : + is_central_scalar 𝕜 (pre_lp E) := +pi.is_central_scalar + +variables [∀ i, has_bounded_smul 𝕜 (E i)] [∀ i, has_bounded_smul 𝕜' (E i)] + lemma mem_lp_const_smul (c : 𝕜) (f : lp E p) : c • (f : pre_lp E) ∈ lp E p := (lp.mem_ℓp f).const_smul c @@ -588,43 +605,74 @@ instance : module 𝕜 (lp E p) := @[simp] lemma coe_fn_smul (c : 𝕜) (f : lp E p) : ⇑(c • f) = c • f := rfl -lemma norm_const_smul (hp : p ≠ 0) {c : 𝕜} (f : lp E p) : ‖c • f‖ = ‖c‖ * ‖f‖ := +instance [Π i, smul_comm_class 𝕜' 𝕜 (E i)] : smul_comm_class 𝕜' 𝕜 (lp E p) := +⟨λ r c f, subtype.ext $ smul_comm _ _ _⟩ + +instance [has_smul 𝕜' 𝕜] [Π i, is_scalar_tower 𝕜' 𝕜 (E i)] : is_scalar_tower 𝕜' 𝕜 (lp E p) := +⟨λ r c f, subtype.ext $ smul_assoc _ _ _⟩ + +instance [Π i, module 𝕜ᵐᵒᵖ (E i)] [Π i, is_central_scalar 𝕜 (E i)] : + is_central_scalar 𝕜 (lp E p) := +⟨λ r f, subtype.ext $ op_smul_eq_smul _ _⟩ + +lemma norm_const_smul_le (hp : p ≠ 0) (c : 𝕜) (f : lp E p) : ‖c • f‖ ≤ ‖c‖ * ‖f‖ := begin rcases p.trichotomy with rfl | rfl | hp, { exact absurd rfl hp }, { cases is_empty_or_nonempty α; resetI, { simp [lp.eq_zero' f], }, - apply (lp.is_lub_norm (c • f)).unique, - convert (lp.is_lub_norm f).mul_left (norm_nonneg c), - ext a, - simp [coe_fn_smul, norm_smul] }, - { suffices : ‖c • f‖ ^ p.to_real = (‖c‖ * ‖f‖) ^ p.to_real, - { refine real.rpow_left_inj_on hp.ne' _ _ this, - { exact norm_nonneg' _ }, - { exact mul_nonneg (norm_nonneg _) (norm_nonneg' _) } }, - apply (lp.has_sum_norm hp (c • f)).unique, - convert (lp.has_sum_norm hp f).mul_left (‖c‖ ^ p.to_real), - { simp [coe_fn_smul, norm_smul, real.mul_rpow (norm_nonneg c) (norm_nonneg _)] }, - have hf : 0 ≤ ‖f‖ := lp.norm_nonneg' f, - simp [coe_fn_smul, norm_smul, real.mul_rpow (norm_nonneg c) hf] } + have hcf := lp.is_lub_norm (c • f), + have hfc := (lp.is_lub_norm f).mul_left (norm_nonneg c), + simp_rw [←set.range_comp, function.comp] at hfc, + -- TODO: some `is_lub` API should make it a one-liner from here. + refine hcf.right _, + have := hfc.left, + simp_rw [mem_upper_bounds, set.mem_range, forall_exists_index, + forall_apply_eq_imp_iff'] at this ⊢, + intro a, + exact (norm_smul_le _ _).trans (this a) }, + { letI inst : has_nnnorm (lp E p) := ⟨λ f, ⟨‖f‖, norm_nonneg' _⟩⟩, + have coe_nnnorm : ∀ f : lp E p, ↑‖f‖₊ = ‖f‖ := λ _, rfl, + suffices : ‖c • f‖₊ ^ p.to_real ≤ (‖c‖₊ * ‖f‖₊) ^ p.to_real, + { rwa nnreal.rpow_le_rpow_iff hp at this }, + unfreezingI { clear_value inst }, + rw [nnreal.mul_rpow], + have hLHS := (lp.has_sum_norm hp (c • f)), + have hRHS := (lp.has_sum_norm hp f).mul_left (‖c‖ ^ p.to_real), + simp_rw [←coe_nnnorm, ←_root_.coe_nnnorm, ←nnreal.coe_rpow, ←nnreal.coe_mul, + nnreal.has_sum_coe] at hRHS hLHS, + refine has_sum_mono hLHS hRHS (λ i, _), + dsimp only, + rw [←nnreal.mul_rpow], + exact nnreal.rpow_le_rpow (nnnorm_smul_le _ _) ennreal.to_real_nonneg } end -instance [fact (1 ≤ p)] : normed_space 𝕜 (lp E p) := -{ norm_smul_le := λ c f, begin - have hp : 0 < p := zero_lt_one.trans_le (fact.out _), - simp [norm_const_smul hp.ne'] - end } +instance [fact (1 ≤ p)] : has_bounded_smul 𝕜 (lp E p) := +has_bounded_smul.of_norm_smul_le $ norm_const_smul_le (zero_lt_one.trans_le $ fact.out (1 ≤ p)).ne' + +end has_bounded_smul -variables {𝕜' : Type*} [normed_field 𝕜'] +section division_ring +variables {𝕜 : Type*} +variables [normed_division_ring 𝕜] [Π i, module 𝕜 (E i)] [∀ i, has_bounded_smul 𝕜 (E i)] -instance [Π i, normed_space 𝕜' (E i)] [has_smul 𝕜' 𝕜] [Π i, is_scalar_tower 𝕜' 𝕜 (E i)] : - is_scalar_tower 𝕜' 𝕜 (lp E p) := +lemma norm_const_smul (hp : p ≠ 0) {c : 𝕜} (f : lp E p) : ‖c • f‖ = ‖c‖ * ‖f‖ := begin - refine ⟨λ r c f, _⟩, - ext1, - exact (lp.coe_fn_smul _ _).trans (smul_assoc _ _ _) + obtain rfl | hc := eq_or_ne c 0, + { simp }, + refine le_antisymm (norm_const_smul_le hp c f) _, + have := mul_le_mul_of_nonneg_left (norm_const_smul_le hp c⁻¹ (c • f)) (norm_nonneg c), + rwa [inv_smul_smul₀ hc, norm_inv, mul_inv_cancel_left₀ (norm_ne_zero_iff.mpr hc)] at this, end +end division_ring + +section normed_space +variables {𝕜 : Type*} [normed_field 𝕜] [Π i, normed_space 𝕜 (E i)] + +instance [fact (1 ≤ p)] : normed_space 𝕜 (lp E p) := +{ norm_smul_le := λ c f, norm_smul_le _ _} + end normed_space section normed_star_group @@ -667,8 +715,8 @@ instance [hp : fact (1 ≤ p)] : normed_star_group (lp E p) := { simp only [lp.norm_eq_tsum_rpow h, lp.star_apply, norm_star] } end } -variables {𝕜 : Type*} [has_star 𝕜] [normed_field 𝕜] -variables [Π i, normed_space 𝕜 (E i)] [Π i, star_module 𝕜 (E i)] +variables {𝕜 : Type*} [has_star 𝕜] [normed_ring 𝕜] +variables [Π i, module 𝕜 (E i)] [∀ i, has_bounded_smul 𝕜 (E i)] [Π i, star_module 𝕜 (E i)] instance : star_module 𝕜 (lp E p) := { star_smul := λ r f, ext $ star_smul _ _ } @@ -710,12 +758,14 @@ instance : non_unital_normed_ring (lp B ∞) := -- we also want a `non_unital_normed_comm_ring` instance, but this has to wait for #13719 -instance infty_is_scalar_tower {𝕜} [normed_field 𝕜] [Π i, normed_space 𝕜 (B i)] +instance infty_is_scalar_tower + {𝕜} [normed_ring 𝕜] [Π i, module 𝕜 (B i)] [∀ i, has_bounded_smul 𝕜 (B i)] [Π i, is_scalar_tower 𝕜 (B i) (B i)] : is_scalar_tower 𝕜 (lp B ∞) (lp B ∞) := ⟨λ r f g, lp.ext $ smul_assoc r ⇑f ⇑g⟩ -instance infty_smul_comm_class {𝕜} [normed_field 𝕜] [Π i, normed_space 𝕜 (B i)] +instance infty_smul_comm_class + {𝕜} [normed_ring 𝕜] [Π i, module 𝕜 (B i)] [∀ i, has_bounded_smul 𝕜 (B i)] [Π i, smul_comm_class 𝕜 (B i) (B i)] : smul_comm_class 𝕜 (lp B ∞) (lp B ∞) := ⟨λ r f g, lp.ext $ smul_comm r ⇑f ⇑g⟩ @@ -847,7 +897,7 @@ instance infty_normed_algebra : normed_algebra 𝕜 (lp B ∞) := end algebra section single -variables {𝕜 : Type*} [normed_field 𝕜] [Π i, normed_space 𝕜 (E i)] +variables {𝕜 : Type*} [normed_ring 𝕜] [Π i, module 𝕜 (E i)] [∀ i, has_bounded_smul 𝕜 (E i)] variables [decidable_eq α] /-- The element of `lp E p` which is `a : E i` at the index `i`, and zero elsewhere. -/ From 5a684ce82399d820475609907c6ef8dba5b1b97c Mon Sep 17 00:00:00 2001 From: Emily Witt Date: Thu, 25 May 2023 19:59:17 +0000 Subject: [PATCH 1078/1291] feat(algebra/homology/local_cohomology): just the definition (#19061) Co-authored-by: Emily Witt Co-authored-by: Jake Levinson Co-authored-by: Scott Morrison --- docs/references.bib | 50 ++++++ src/algebra/category/Module/colimits.lean | 6 + src/algebra/homology/local_cohomology.lean | 189 +++++++++++++++++++++ 3 files changed, 245 insertions(+) create mode 100644 src/algebra/homology/local_cohomology.lean diff --git a/docs/references.bib b/docs/references.bib index d5c6b40fab9e3..51d5318f01f37 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -386,6 +386,20 @@ @Book{ boydVandenberghe2004 url = {https://web.stanford.edu/~boyd/cvxbook/bv_cvxbook.pdf} } +@Book{ brodmannsharp13, + author = {Brodmann, M. P. and Sharp, R. Y.}, + title = {Local cohomology}, + series = {Cambridge Studies in Advanced Mathematics}, + volume = {136}, + edition = {Second}, + note = {An algebraic introduction with geometric applications}, + publisher = {Cambridge University Press, Cambridge}, + year = {2013}, + pages = {xxii+491}, + isbn = {978-0-521-51363-0}, + mrclass = {13D45 (13-01)} +} + @Book{ cabreragarciarodriguezpalacios2014, author = {Miguel {Cabrera Garc\'{\i}a} and \'Angel {Rodr\'{\i}guez Palacios}}, @@ -1195,6 +1209,20 @@ @Book{ harmandwernerwerner1993 zbl = {0789.46011} } +@Book{ hartshorne61, + author = {Hartshorne, Robin}, + title = {Local cohomology}, + series = {Lecture Notes in Mathematics, No. 41}, + note = {A seminar given by A. Grothendieck, Harvard University, + Fall, 1961}, + publisher = {Springer-Verlag, Berlin-New York}, + year = {1967}, + pages = {vi+106}, + mrclass = {14.55 (18.00)}, + mrnumber = {0224620}, + mrreviewer = {F. Oort} +} + @Article{ Haze09, title = {Witt vectors. Part 1}, isbn = {9780444532572}, @@ -1219,6 +1247,12 @@ @Article{ Higman52 year = {1952} } +@Unpublished{ hochsterunpublished, + title = {Local cohomology}, + author = {Hochster, Mel}, + url = {https://dept.math.lsa.umich.edu/~hochster/615W11/loc.pdf} +} + @Book{ Hodges97, author = {Hodges, Wilfrid}, title = {A Shorter Model Theory}, @@ -1299,6 +1333,22 @@ @Book{ IrelandRosen1990 url = {https://doi.org/10.1007/978-1-4757-2103-4} } +@Book{ iyengaretal07, + author = {Iyengar, Srikanth B. and Leuschke, Graham J. and Leykin, + Anton and Miller, Claudia and Miller, Ezra and Singh, + Anurag K. and Walther, Uli}, + title = {Twenty-four hours of local cohomology}, + series = {Graduate Studies in Mathematics}, + volume = {87}, + publisher = {American Mathematical Society, Providence, RI}, + year = {2007}, + pages = {xviii+282}, + isbn = {978-0-8218-4126-6}, + mrclass = {13D45 (14B15 55N30)}, + doi = {10.1090/gsm/087}, + url = {https://doi-org.www2.lib.ku.edu/10.1090/gsm/087} +} + @Article{ izhakian2016, title = {Supertropical quadratic forms I}, journal = {Journal of Pure and Applied Algebra}, diff --git a/src/algebra/category/Module/colimits.lean b/src/algebra/category/Module/colimits.lean index 33d2030132d88..eaca92284e5c7 100644 --- a/src/algebra/category/Module/colimits.lean +++ b/src/algebra/category/Module/colimits.lean @@ -369,6 +369,12 @@ instance has_colimits_Module : has_colimits (Module.{max v u} R) := { cocone := colimit_cocone F, is_colimit := colimit_cocone_is_colimit F } } } +instance has_colimits_of_size_Module : has_colimits_of_size.{v} (Module.{max v u} R) := +has_colimits_of_size_shrink _ + +instance has_colimits_of_size_zero_Module : has_colimits_of_size.{0} (Module.{max v u} R) := +@has_colimits_of_size_shrink.{0} (Module.{max v u} R) _ Module.colimits.has_colimits_Module + -- We manually add a `has_colimits` instance with universe parameters swapped, for otherwise -- the instance is not found by typeclass search. instance has_colimits_Module' (R : Type u) [ring R] : diff --git a/src/algebra/homology/local_cohomology.lean b/src/algebra/homology/local_cohomology.lean new file mode 100644 index 0000000000000..ec568744815d9 --- /dev/null +++ b/src/algebra/homology/local_cohomology.lean @@ -0,0 +1,189 @@ +/- +Copyright (c) 2023 Emily Witt. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Emily Witt, Scott Morrison, Jake Levinson, Sam van Gool +-/ +import ring_theory.ideal.basic +import algebra.category.Module.colimits +import algebra.category.Module.projective +import category_theory.abelian.ext +import ring_theory.finiteness + +/-! +# Local cohomology. + +This file defines the `i`-th local cohomology module of an `R`-module `M` with support in an +ideal `I` of `R`, where `R` is a commutative ring, as the direct limit of Ext modules: + +Given a collection of ideals cofinal with the powers of `I`, consider the directed system of +quotients of `R` by these ideals, and take the direct limit of the system induced on the `i`-th +Ext into `M`. One can, of course, take the collection to simply be the integral powers of `I`. + +## References + +* [M. Hochster, *Local cohomology*][hochsterunpublished] + +* [R. Hartshorne, *Local cohomology: A seminar given by A. Grothendieck*][hartshorne61] +* [M. Brodmann and R. Sharp, *Local cohomology: An algebraic introduction with geometric + applications*][brodmannsharp13] +* [S. Iyengar, G. Leuschke, A. Leykin, Anton, C. Miller, E. Miller, A. Singh, U. Walther, + *Twenty-four hours of local cohomology*][iyengaretal13] + +## Tags + +local cohomology, local cohomology modules + +## Future work + +* Prove that this definition is equivalent to: + * the right-derived functor definition + * the characterization as the limit of Koszul homology + * the characterization as the cohomology of a Cech-like complex +* Prove that local cohomology depends only on the radical of the ideal +* Establish long exact sequence(s) in local cohomology +-/ + +open opposite +open category_theory +open category_theory.limits + +noncomputable theory + +universes u v + + +namespace local_cohomology + +/- We define local cohomology, implemented as a direct limit of `Ext(R/J, -)`. -/ +section +variables {R : Type u} [comm_ring R] {D : Type v} [small_category D] + +/-- The directed system of `R`-modules of the form `R/J`, where `J` is an ideal of `R`, +determined by the functor `I` -/ +def ring_mod_ideals (I : D ⥤ ideal R) : D ⥤ Module.{u} R := +{ obj := λ t, Module.of R $ R ⧸ (I.obj t), + map := λ s t w, submodule.mapq _ _ (linear_map.id) (I.map w).down.down } + +/- TODO: Once this file is ported, move this file to the right location. -/ +instance Module_enough_projectives' : enough_projectives (Module.{u} R) := + Module.Module_enough_projectives.{u} + +/-- The diagram we will take the colimit of to define local cohomology, corresponding to the +directed system determined by the functor `I` -/ +def diagram (I : D ⥤ ideal R) (i : ℕ) : Dᵒᵖ ⥤ Module.{u} R ⥤ Module.{u} R := +(ring_mod_ideals I).op ⋙ Ext R (Module.{u} R) i + +end + +section +-- We momentarily need to work with a type inequality, as later we will take colimits +-- along diagrams either in Type, or in the same universe as the ring, and we need to cover both. +variables {R : Type max u v} [comm_ring R] {D : Type v} [small_category D] + +/-- `local_cohomology.of_diagram I i` is the the functor sending a module `M` over a commutative +ring `R` to the direct limit of `Ext^i(R/J, M)`, where `J` ranges over a collection of ideals +of `R`, represented as a functor `I`. -/ +/- +In this definition we do not assume any special property of the diagram `I`, but the relevant case +will be where `I` is (cofinal with) the diagram of powers of a single given ideal. + +Below, we give two equivalent (to be shown) definitions of the usual local cohomology with support +in an ideal `J`, `local_cohomology` and `local_cohomology.of_self_le_radical`. + +TODO: Show that any functor cofinal with `I` gives the same result. + -/ +def of_diagram (I : D ⥤ ideal R) (i : ℕ) : + Module.{max u v} R ⥤ Module.{max u v} R := +colimit (diagram.{(max u v) v} I i) + +end + +section diagrams + +variables {R : Type u} [comm_ring R] + +/-- The functor sending a natural number `i` to the `i`-th power of the ideal `J` -/ +def ideal_powers_diagram (J : ideal R) : ℕᵒᵖ ⥤ ideal R := +{ obj := λ t, J^(unop t), + map := λ s t w, ⟨⟨ideal.pow_le_pow w.unop.down.down⟩⟩, } + +/-- The full subcategory of all ideals with radical containing `J` -/ +@[derive category] def self_le_radical (J : ideal R) : Type u := +full_subcategory (λ J' : ideal R, J ≤ J'.radical) + +instance self_le_radical.inhabited (J : ideal R) : inhabited (self_le_radical J) := +{ default := ⟨J, ideal.le_radical⟩ } + +/-- The diagram of all ideals with radical containing `J`, represented as a functor. +This is the "largest" diagram that computes local cohomology with support in `J`. -/ +def self_le_radical_diagram (J : ideal R) : (self_le_radical J) ⥤ ideal R := +full_subcategory_inclusion _ + +end diagrams + +end local_cohomology + +/-! We give two models for the local cohomology with support in an ideal `J`: first in terms of +the powers of `J` (`local_cohomology`), then in terms of *all* ideals with radical +containing `J` (`local_cohomology.of_self_le_radical`). -/ +section models_for_local_cohomology + +open local_cohomology + +variables {R : Type u} [comm_ring R] + +/-- `local_cohomology J i` is `i`-th the local cohomology module of a module `M` over +a commutative ring `R` with support in the ideal `J` of `R`, defined as the direct limit +of `Ext^i(R/J^t, M)` over all powers `t : ℕ`. -/ +def local_cohomology (J : ideal R) (i : ℕ) : Module.{u} R ⥤ Module.{u} R := + of_diagram (ideal_powers_diagram J) i + +/-- Local cohomology as the direct limit of `Ext^i(R/J', M)` over *all* ideals `J'` with radical +containing `J`. -/ +def local_cohomology.of_self_le_radical (J : ideal R) (i : ℕ) : + Module.{u} R ⥤ Module.{u} R := +of_diagram.{u} (self_le_radical_diagram.{u} J) i +/- TODO: Construct `local_cohomology J i ≅ local_cohomology.of_self_le_radical J i`. Use this to +show that local cohomology depends only on `J.radical`. -/ + +end models_for_local_cohomology + + +section local_cohomology_equiv + +open local_cohomology + +variables {R : Type u} [comm_ring R] (I J : ideal R) + +/-- Lifting `ideal_powers_diagram J` from a diagram valued in `ideals R` to a diagram +valued in `self_le_radical J`. -/ +def local_cohomology.ideal_powers_to_self_le_radical (J : ideal R) : ℕᵒᵖ ⥤ self_le_radical J := +full_subcategory.lift _ (ideal_powers_diagram J) +(λ k, begin + change _ ≤ (J^(unop k)).radical, + cases (unop k), + { simp only [ideal.radical_top, pow_zero, ideal.one_eq_top, le_top] }, + { simp only [J.radical_pow _ n.succ_pos, ideal.le_radical] }, +end) + +/-- The composition with the inclusion into `ideals R` is isomorphic to `ideal_powers_diagram J`. -/ +def local_cohomology.ideal_powers_to_self_le_radical_comp_inclusion (J : ideal R) : +local_cohomology.ideal_powers_to_self_le_radical J ⋙ self_le_radical_diagram J + ≅ ideal_powers_diagram J := + full_subcategory.lift_comp_inclusion _ _ _ + +/-- The lemma below essentially says that `ideal_powers_to_self_le_radical I` is initial in +`self_le_radical_diagram I`. + +PORTING NOTE: This lemma should probably be moved to `ring_theory/finiteness.lean` +to be near `ideal.exists_radical_pow_le_of_fg`, which it generalizes. -/ +lemma ideal.exists_pow_le_of_le_radical_of_fg (hIJ : I ≤ J.radical) (hJ : J.radical.fg) : + ∃ (k : ℕ), I^k ≤ J := +begin + obtain ⟨k, hk⟩ := J.exists_radical_pow_le_of_fg hJ, + use k, + calc I^k ≤ J.radical^k : ideal.pow_mono hIJ _ + ... ≤ J : hk, +end + +end local_cohomology_equiv From e1a18cad9cd462973d760af7de36b05776b8811c Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 25 May 2023 19:59:18 +0000 Subject: [PATCH 1079/1291] =?UTF-8?q?feat(analysis/special=5Ffunctions/exp?= =?UTF-8?q?onential):=20derivative=20of=20`u=20=E2=86=A6=20exp=20?= =?UTF-8?q?=F0=9D=95=82=20(u=20=E2=80=A2=20x)`=20(#19062)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Revived from an old branch of @ADedecker, golfed using new lemmas that I added in the meantime. --- .../special_functions/exponential.lean | 202 +++++++++++++++++- 1 file changed, 199 insertions(+), 3 deletions(-) diff --git a/src/analysis/special_functions/exponential.lean b/src/analysis/special_functions/exponential.lean index dccb029c003ca..3208685d55087 100644 --- a/src/analysis/special_functions/exponential.lean +++ b/src/analysis/special_functions/exponential.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2021 Anatole Dedecker. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Anatole Dedecker +Authors: Anatole Dedecker, Eric Wieser -/ import analysis.normed_space.exponential import analysis.calculus.fderiv_analytic @@ -24,17 +24,24 @@ We prove most result for an arbitrary field `𝕂`, and then specialize to `𝕂 `1 : 𝔸 →L[𝕂] 𝔸` at zero, as long as it converges on a neighborhood of zero (see also `has_strict_deriv_at_exp_zero_of_radius_pos` for the case `𝔸 = 𝕂`) - `has_strict_fderiv_at_exp_of_lt_radius` : if `𝕂` has characteristic zero and `𝔸` is commutative, - then given a point `x` in the disk of convergence, `exp 𝕂` as strict Fréchet-derivative + then given a point `x` in the disk of convergence, `exp 𝕂` has strict Fréchet-derivative `exp 𝕂 x • 1 : 𝔸 →L[𝕂] 𝔸` at x (see also `has_strict_deriv_at_exp_of_lt_radius` for the case `𝔸 = 𝕂`) +- `has_strict_fderiv_at_exp_smul_const_of_mem_ball`: even when `𝔸` is non-commutative, if we have + an intermediate algebra `𝕊` which is commutative, then the function `(u : 𝕊) ↦ exp 𝕂 (u • x)`, + still has strict Fréchet-derivative `exp 𝕂 (t • x) • (1 : 𝕊 →L[𝕂] 𝕊).smul_right x` at `t` if + `t • x` is in the radius of convergence. ### `𝕂 = ℝ` or `𝕂 = ℂ` - `has_strict_fderiv_at_exp_zero` : `exp 𝕂` has strict Fréchet-derivative `1 : 𝔸 →L[𝕂] 𝔸` at zero (see also `has_strict_deriv_at_exp_zero` for the case `𝔸 = 𝕂`) -- `has_strict_fderiv_at_exp` : if `𝔸` is commutative, then given any point `x`, `exp 𝕂` as strict +- `has_strict_fderiv_at_exp` : if `𝔸` is commutative, then given any point `x`, `exp 𝕂` has strict Fréchet-derivative `exp 𝕂 x • 1 : 𝔸 →L[𝕂] 𝔸` at x (see also `has_strict_deriv_at_exp` for the case `𝔸 = 𝕂`) +- `has_strict_fderiv_at_exp_smul_const`: even when `𝔸` is non-commutative, if we have + an intermediate algebra `𝕊` which is commutative, then the function `(u : 𝕊) ↦ exp 𝕂 (u • x)` + still has strict Fréchet-derivative `exp 𝕂 (t • x) • (1 : 𝔸 →L[𝕂] 𝔸).smul_right x` at `t`. ### Compatibilty with `real.exp` and `complex.exp` @@ -211,3 +218,192 @@ end lemma real.exp_eq_exp_ℝ : real.exp = exp ℝ := by { ext x, exact_mod_cast congr_fun complex.exp_eq_exp_ℂ x } + +/-! ### Derivative of $\exp (ux)$ by $u$ + +Note that since for `x : 𝔸` we have `normed_ring 𝔸` not `normed_comm_ring 𝔸`, we cannot deduce +these results from `has_fderiv_at_exp_of_mem_ball` applied to the algebra `𝔸`. + +One possible solution for that would be to apply `has_fderiv_at_exp_of_mem_ball` to the +commutative algebra `algebra.elemental_algebra 𝕊 x`. Unfortunately we don't have all the required +API, so we leave that to a future refactor (see leanprover-community/mathlib#19062 for discussion). + +We could also go the other way around and deduce `has_fderiv_at_exp_of_mem_ball` from +`has_fderiv_at_exp_smul_const_of_mem_ball` applied to `𝕊 := 𝔸`, `x := (1 : 𝔸)`, and `t := x`. +However, doing so would make the aformentioned `elemental_algebra` refactor harder, so for now we +just prove these two lemmas independently. + +A last strategy would be to deduce everything from the more general non-commutative case, +$$\frac{d}{dt}e^{x(t)} = \int_0^1 e^{sx(t)} \left(\frac{d}{dt}e^{x(t)}\right) e^{(1-s)x(t)} ds$$ +but this is harder to prove, and typically is shown by going via these results first. + +TODO: prove this result too! +-/ + +section exp_smul +variables {𝕂 𝕊 𝔸 : Type*} +variables (𝕂) + +open_locale topology +open asymptotics filter + +section mem_ball +variables [nontrivially_normed_field 𝕂] [char_zero 𝕂] +variables [normed_comm_ring 𝕊] [normed_ring 𝔸] +variables [normed_space 𝕂 𝕊] [normed_algebra 𝕂 𝔸] [algebra 𝕊 𝔸] [has_continuous_smul 𝕊 𝔸] +variables [is_scalar_tower 𝕂 𝕊 𝔸] +variables [complete_space 𝔸] + +lemma has_fderiv_at_exp_smul_const_of_mem_ball + (x : 𝔸) (t : 𝕊) (htx : t • x ∈ emetric.ball (0 : 𝔸) (exp_series 𝕂 𝔸).radius) : + has_fderiv_at (λ u : 𝕊, exp 𝕂 (u • x)) + (exp 𝕂 (t • x) • (1 : 𝕊 →L[𝕂] 𝕊).smul_right x) t := +begin + -- TODO: prove this via `has_fderiv_at_exp_of_mem_ball` using the commutative ring + -- `algebra.elemental_algebra 𝕊 x`. See leanprover-community/mathlib#19062 for discussion. + have hpos : 0 < (exp_series 𝕂 𝔸).radius := (zero_le _).trans_lt htx, + rw has_fderiv_at_iff_is_o_nhds_zero, + suffices : + (λ h, exp 𝕂 (t • x) * (exp 𝕂 ((0 + h) • x) - exp 𝕂 ((0 : 𝕊) • x) + - ((1 : 𝕊 →L[𝕂] 𝕊).smul_right x) h)) + =ᶠ[𝓝 0] (λ h, exp 𝕂 ((t + h) • x) - exp 𝕂 (t • x) + - (exp 𝕂 (t • x) • (1 : 𝕊 →L[𝕂] 𝕊).smul_right x) h), + { refine (is_o.const_mul_left _ _).congr' this (eventually_eq.refl _ _), + rw ← @has_fderiv_at_iff_is_o_nhds_zero _ _ _ _ _ _ _ _ + (λ u, exp 𝕂 (u • x)) ((1 : 𝕊 →L[𝕂] 𝕊).smul_right x) 0, + have : has_fderiv_at (exp 𝕂) (1 : 𝔸 →L[𝕂] 𝔸) ((1 : 𝕊 →L[𝕂] 𝕊).smul_right x 0), + { rw [continuous_linear_map.smul_right_apply, continuous_linear_map.one_apply, zero_smul], + exact has_fderiv_at_exp_zero_of_radius_pos hpos }, + exact this.comp 0 ((1 : 𝕊 →L[𝕂] 𝕊).smul_right x).has_fderiv_at }, + have : tendsto (λ h : 𝕊, h • x) (𝓝 0) (𝓝 0), + { rw ← zero_smul 𝕊 x, + exact tendsto_id.smul_const x }, + have : ∀ᶠ h in 𝓝 (0 : 𝕊), h • x ∈ emetric.ball (0 : 𝔸) (exp_series 𝕂 𝔸).radius := + this.eventually (emetric.ball_mem_nhds _ hpos), + filter_upwards [this], + intros h hh, + have : commute (t • x) (h • x) := ((commute.refl x).smul_left t).smul_right h, + rw [add_smul t h, exp_add_of_commute_of_mem_ball this htx hh, zero_add, zero_smul, exp_zero, + continuous_linear_map.smul_right_apply, continuous_linear_map.one_apply, + continuous_linear_map.smul_apply, continuous_linear_map.smul_right_apply, + continuous_linear_map.one_apply, smul_eq_mul, mul_sub_left_distrib, mul_sub_left_distrib, + mul_one], +end + +lemma has_fderiv_at_exp_smul_const_of_mem_ball' + (x : 𝔸) (t : 𝕊) (htx : t • x ∈ emetric.ball (0 : 𝔸) (exp_series 𝕂 𝔸).radius) : + has_fderiv_at (λ u : 𝕊, exp 𝕂 (u • x)) + (((1 : 𝕊 →L[𝕂] 𝕊).smul_right x).smul_right (exp 𝕂 (t • x))) t := +begin + convert has_fderiv_at_exp_smul_const_of_mem_ball 𝕂 _ _ htx using 1, + ext t', + show commute (t' • x) (exp 𝕂 (t • x)), + exact (((commute.refl x).smul_left t').smul_right t).exp_right 𝕂, +end + +lemma has_strict_fderiv_at_exp_smul_const_of_mem_ball (x : 𝔸) (t : 𝕊) + (htx : t • x ∈ emetric.ball (0 : 𝔸) (exp_series 𝕂 𝔸).radius) : + has_strict_fderiv_at (λ u : 𝕊, exp 𝕂 (u • x)) + (exp 𝕂 (t • x) • (1 : 𝕊 →L[𝕂] 𝕊).smul_right x) t := +let ⟨p, hp⟩ := analytic_at_exp_of_mem_ball (t • x) htx in +have deriv₁ : has_strict_fderiv_at (λ u : 𝕊, exp 𝕂 (u • x)) _ t, + from hp.has_strict_fderiv_at.comp t + ((continuous_linear_map.id 𝕂 𝕊).smul_right x).has_strict_fderiv_at, +have deriv₂ : has_fderiv_at (λ u : 𝕊, exp 𝕂 (u • x)) _ t, + from has_fderiv_at_exp_smul_const_of_mem_ball 𝕂 x t htx, +(deriv₁.has_fderiv_at.unique deriv₂) ▸ deriv₁ + +lemma has_strict_fderiv_at_exp_smul_const_of_mem_ball' (x : 𝔸) (t : 𝕊) + (htx : t • x ∈ emetric.ball (0 : 𝔸) (exp_series 𝕂 𝔸).radius) : + has_strict_fderiv_at (λ u : 𝕊, exp 𝕂 (u • x)) + (((1 : 𝕊 →L[𝕂] 𝕊).smul_right x).smul_right (exp 𝕂 (t • x))) t := +let ⟨p, hp⟩ := analytic_at_exp_of_mem_ball (t • x) htx in +begin + convert has_strict_fderiv_at_exp_smul_const_of_mem_ball 𝕂 _ _ htx using 1, + ext t', + show commute (t' • x) (exp 𝕂 (t • x)), + exact (((commute.refl x).smul_left t').smul_right t).exp_right 𝕂, +end + +variables {𝕂} + +lemma has_strict_deriv_at_exp_smul_const_of_mem_ball (x : 𝔸) (t : 𝕂) + (htx : t • x ∈ emetric.ball (0 : 𝔸) (exp_series 𝕂 𝔸).radius) : + has_strict_deriv_at (λ u : 𝕂, exp 𝕂 (u • x)) (exp 𝕂 (t • x) * x) t := +by simpa using (has_strict_fderiv_at_exp_smul_const_of_mem_ball 𝕂 x t htx).has_strict_deriv_at + +lemma has_strict_deriv_at_exp_smul_const_of_mem_ball' (x : 𝔸) (t : 𝕂) + (htx : t • x ∈ emetric.ball (0 : 𝔸) (exp_series 𝕂 𝔸).radius) : + has_strict_deriv_at (λ u : 𝕂, exp 𝕂 (u • x)) (x * exp 𝕂 (t • x)) t := +by simpa using (has_strict_fderiv_at_exp_smul_const_of_mem_ball' 𝕂 x t htx).has_strict_deriv_at + +lemma has_deriv_at_exp_smul_const_of_mem_ball (x : 𝔸) (t : 𝕂) + (htx : t • x ∈ emetric.ball (0 : 𝔸) (exp_series 𝕂 𝔸).radius) : + has_deriv_at (λ u : 𝕂, exp 𝕂 (u • x)) (exp 𝕂 (t • x) * x) t := +(has_strict_deriv_at_exp_smul_const_of_mem_ball x t htx).has_deriv_at + +lemma has_deriv_at_exp_smul_const_of_mem_ball' (x : 𝔸) (t : 𝕂) + (htx : t • x ∈ emetric.ball (0 : 𝔸) (exp_series 𝕂 𝔸).radius) : + has_deriv_at (λ u : 𝕂, exp 𝕂 (u • x)) (x * exp 𝕂 (t • x)) t := +(has_strict_deriv_at_exp_smul_const_of_mem_ball' x t htx).has_deriv_at + +end mem_ball + +section is_R_or_C +variables [is_R_or_C 𝕂] +variables [normed_comm_ring 𝕊] [normed_ring 𝔸] +variables [normed_algebra 𝕂 𝕊] [normed_algebra 𝕂 𝔸] [algebra 𝕊 𝔸] [has_continuous_smul 𝕊 𝔸] +variables [is_scalar_tower 𝕂 𝕊 𝔸] +variables [complete_space 𝔸] + +variables (𝕂) + +lemma has_fderiv_at_exp_smul_const (x : 𝔸) (t : 𝕊) : + has_fderiv_at (λ u : 𝕊, exp 𝕂 (u • x)) + (exp 𝕂 (t • x) • (1 : 𝕊 →L[𝕂] 𝕊).smul_right x) t := +has_fderiv_at_exp_smul_const_of_mem_ball 𝕂 _ _ $ + (exp_series_radius_eq_top 𝕂 𝔸).symm ▸ edist_lt_top _ _ + +lemma has_fderiv_at_exp_smul_const' (x : 𝔸) (t : 𝕊) : + has_fderiv_at (λ u : 𝕊, exp 𝕂 (u • x)) + (((1 : 𝕊 →L[𝕂] 𝕊).smul_right x).smul_right (exp 𝕂 (t • x))) t := +has_fderiv_at_exp_smul_const_of_mem_ball' 𝕂 _ _ $ + (exp_series_radius_eq_top 𝕂 𝔸).symm ▸ edist_lt_top _ _ + +lemma has_strict_fderiv_at_exp_smul_const (x : 𝔸) (t : 𝕊) : + has_strict_fderiv_at (λ u : 𝕊, exp 𝕂 (u • x)) + (exp 𝕂 (t • x) • (1 : 𝕊 →L[𝕂] 𝕊).smul_right x) t := +has_strict_fderiv_at_exp_smul_const_of_mem_ball 𝕂 _ _ $ + (exp_series_radius_eq_top 𝕂 𝔸).symm ▸ edist_lt_top _ _ + +lemma has_strict_fderiv_at_exp_smul_const' (x : 𝔸) (t : 𝕊) : + has_strict_fderiv_at (λ u : 𝕊, exp 𝕂 (u • x)) + (((1 : 𝕊 →L[𝕂] 𝕊).smul_right x).smul_right (exp 𝕂 (t • x))) t := +has_strict_fderiv_at_exp_smul_const_of_mem_ball' 𝕂 _ _ $ + (exp_series_radius_eq_top 𝕂 𝔸).symm ▸ edist_lt_top _ _ + +variables {𝕂} + +lemma has_strict_deriv_at_exp_smul_const (x : 𝔸) (t : 𝕂) : + has_strict_deriv_at (λ u : 𝕂, exp 𝕂 (u • x)) (exp 𝕂 (t • x) * x) t := +has_strict_deriv_at_exp_smul_const_of_mem_ball _ _ $ + (exp_series_radius_eq_top 𝕂 𝔸).symm ▸ edist_lt_top _ _ + +lemma has_strict_deriv_at_exp_smul_const' (x : 𝔸) (t : 𝕂) : + has_strict_deriv_at (λ u : 𝕂, exp 𝕂 (u • x)) (x * exp 𝕂 (t • x)) t := +has_strict_deriv_at_exp_smul_const_of_mem_ball' _ _ $ + (exp_series_radius_eq_top 𝕂 𝔸).symm ▸ edist_lt_top _ _ + +lemma has_deriv_at_exp_smul_const (x : 𝔸) (t : 𝕂) : + has_deriv_at (λ u : 𝕂, exp 𝕂 (u • x)) (exp 𝕂 (t • x) * x) t := +has_deriv_at_exp_smul_const_of_mem_ball _ _ $ + (exp_series_radius_eq_top 𝕂 𝔸).symm ▸ edist_lt_top _ _ + +lemma has_deriv_at_exp_smul_const' (x : 𝔸) (t : 𝕂) : + has_deriv_at (λ u : 𝕂, exp 𝕂 (u • x)) (x * exp 𝕂 (t • x)) t := +has_deriv_at_exp_smul_const_of_mem_ball' _ _ $ + (exp_series_radius_eq_top 𝕂 𝔸).symm ▸ edist_lt_top _ _ + +end is_R_or_C + +end exp_smul From 45a46f4f03f8ae41491bf3605e8e0e363ba192fd Mon Sep 17 00:00:00 2001 From: Eric Rodriguez <37984851+ericrbg@users.noreply.github.com> Date: Thu, 25 May 2023 21:34:40 +0000 Subject: [PATCH 1080/1291] doc(analysis/complex/arg): fix docs (#19095) This is being backported from leanprover-community/mathlib4#4355. --- src/analysis/complex/arg.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/analysis/complex/arg.lean b/src/analysis/complex/arg.lean index a1716af4733cb..cdebb9f3deb68 100644 --- a/src/analysis/complex/arg.lean +++ b/src/analysis/complex/arg.lean @@ -17,8 +17,8 @@ the usual way this is considered. * `complex.same_ray_iff` : Two complex numbers are on the same ray iff one of them is zero, or they have the same argument. -* `complex.abs_add_eq/complex.abs_sub_eq`: If two non zero complex numbers have different argument, - then the triangle inequality becomes strict. +* `complex.abs_add_eq/complex.abs_sub_eq`: If two non zero complex numbers have the same argument, + then the triangle inequality is an equality. -/ From d91e7f7a7f1c7e9f0e18fdb6bde4f652004c735d Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Thu, 25 May 2023 22:47:33 +0000 Subject: [PATCH 1081/1291] feat(topology/{maps,separation}): add lemmas about closed and quotient maps (#19071) Lemma statements are from Shamrock-Frost/BrouwerFixedPoint Co-authored-by: @Shamrock-Frost --- src/topology/maps.lean | 14 ++++++++++++-- src/topology/separation.lean | 11 +++++++++-- 2 files changed, 21 insertions(+), 4 deletions(-) diff --git a/src/topology/maps.lean b/src/topology/maps.lean index 1f6641fd51246..cbd80ed6314c7 100644 --- a/src/topology/maps.lean +++ b/src/topology/maps.lean @@ -228,10 +228,15 @@ def quotient_map {α : Type*} {β : Type*} [tα : topological_space α] [tβ : t (f : α → β) : Prop := surjective f ∧ tβ = tα.coinduced f -lemma quotient_map_iff {α β : Type*} [topological_space α] [topological_space β] {f : α → β} : +lemma quotient_map_iff [topological_space α] [topological_space β] {f : α → β} : quotient_map f ↔ surjective f ∧ ∀ s : set β, is_open s ↔ is_open (f ⁻¹' s) := and_congr iff.rfl topological_space_eq_iff +lemma quotient_map_iff_closed [topological_space α] [topological_space β] {f : α → β} : + quotient_map f ↔ surjective f ∧ ∀ s : set β, is_closed s ↔ is_closed (f ⁻¹' s) := +quotient_map_iff.trans $ iff.rfl.and $ compl_surjective.forall.trans $ + by simp only [is_open_compl_iff, preimage_compl] + namespace quotient_map variables [topological_space α] [topological_space β] [topological_space γ] [topological_space δ] @@ -270,7 +275,7 @@ protected lemma is_open_preimage (hf : quotient_map f) {s : set β} : protected lemma is_closed_preimage (hf : quotient_map f) {s : set β} : is_closed (f ⁻¹' s) ↔ is_closed s := -by simp only [← is_open_compl_iff, ← preimage_compl, hf.is_open_preimage] +((quotient_map_iff_closed.1 hf).2 s).symm end quotient_map @@ -427,6 +432,11 @@ end lemma closed_range {f : α → β} (hf : is_closed_map f) : is_closed (range f) := @image_univ _ _ f ▸ hf _ is_closed_univ +lemma to_quotient_map {f : α → β} (hcl : is_closed_map f) (hcont : continuous f) + (hsurj : surjective f) : quotient_map f := +quotient_map_iff_closed.2 + ⟨hsurj, λ s, ⟨λ hs, hs.preimage hcont, λ hs, hsurj.image_preimage s ▸ hcl _ hs⟩⟩ + end is_closed_map lemma inducing.is_closed_map [topological_space α] [topological_space β] diff --git a/src/topology/separation.lean b/src/topology/separation.lean index 95001a4d61cde..c5fc7532ce7e7 100644 --- a/src/topology/separation.lean +++ b/src/topology/separation.lean @@ -1208,14 +1208,21 @@ begin by rw [← diff_inter, hO.inter_eq, diff_empty]⟩ end -lemma continuous.is_closed_map [compact_space α] [t2_space β] {f : α → β} (h : continuous f) : - is_closed_map f := +/-- A continuous map from a compact space to a Hausdorff space is a closed map. -/ +protected lemma continuous.is_closed_map [compact_space α] [t2_space β] {f : α → β} + (h : continuous f) : is_closed_map f := λ s hs, (hs.is_compact.image h).is_closed +/-- An injective continuous map from a compact space to a Hausdorff space is a closed embedding. -/ lemma continuous.closed_embedding [compact_space α] [t2_space β] {f : α → β} (h : continuous f) (hf : function.injective f) : closed_embedding f := closed_embedding_of_continuous_injective_closed h hf h.is_closed_map +/-- A surjective continuous map from a compact space to a Hausdorff space is a quotient map. -/ +lemma quotient_map.of_surjective_continuous [compact_space α] [t2_space β] {f : α → β} + (hsurj : surjective f) (hcont : continuous f) : quotient_map f := +hcont.is_closed_map.to_quotient_map hcont hsurj + section open finset function /-- For every finite open cover `Uᵢ` of a compact set, there exists a compact cover `Kᵢ ⊆ Uᵢ`. -/ From cb42593171ba005beaaf4549fcfe0dece9ada4c9 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Fri, 26 May 2023 03:27:28 +0000 Subject: [PATCH 1082/1291] feat(data/real/basic): add `real.supr_nonneg` etc (#19096) Motivated by lemmas from the sphere eversion project --- src/data/real/basic.lean | 29 ++++++++++++++++++++++++----- 1 file changed, 24 insertions(+), 5 deletions(-) diff --git a/src/data/real/basic.lean b/src/data/real/basic.lean index 4009cb9258435..dbddb697e0b87 100644 --- a/src/data/real/basic.lean +++ b/src/data/real/basic.lean @@ -600,7 +600,7 @@ lemma infi_of_not_bdd_below {α : Sort*} {f : α → ℝ} (hf : ¬ bdd_below (s /-- As `0` is the default value for `real.Sup` of the empty set or sets which are not bounded above, it -suffices to show that `S` is bounded below by `0` to show that `0 ≤ Inf S`. +suffices to show that `S` is bounded below by `0` to show that `0 ≤ Sup S`. -/ lemma Sup_nonneg (S : set ℝ) (hS : ∀ x ∈ S, (0:ℝ) ≤ x) : 0 ≤ Sup S := begin @@ -610,15 +610,34 @@ begin end /-- -As `0` is the default value for `real.Sup` of the empty set, it suffices to show that `S` is -bounded above by `0` to show that `Sup S ≤ 0`. +As `0` is the default value for `real.Sup` of the empty set or sets which are not bounded above, it +suffices to show that `f i` is nonnegative to show that `0 ≤ ⨆ i, f i`. -/ -lemma Sup_nonpos (S : set ℝ) (hS : ∀ x ∈ S, x ≤ (0:ℝ)) : Sup S ≤ 0 := +protected lemma supr_nonneg {ι : Sort*} {f : ι → ℝ} (hf : ∀ i, 0 ≤ f i) : 0 ≤ ⨆ i, f i := +Sup_nonneg _ $ set.forall_range_iff.2 hf + +/-- +As `0` is the default value for `real.Sup` of the empty set or sets which are not bounded above, it +suffices to show that all elements of `S` are bounded by a nonnagative number to show that `Sup S` +is bounded by this number. +-/ +protected lemma Sup_le {S : set ℝ} {a : ℝ} (hS : ∀ x ∈ S, x ≤ a) (ha : 0 ≤ a) : Sup S ≤ a := begin rcases S.eq_empty_or_nonempty with rfl | hS₂, - exacts [Sup_empty.le, cSup_le hS₂ hS], + exacts [Sup_empty.trans_le ha, cSup_le hS₂ hS], end +protected lemma supr_le {ι : Sort*} {f : ι → ℝ} {a : ℝ} (hS : ∀ i, f i ≤ a) (ha : 0 ≤ a) : + (⨆ i, f i) ≤ a := +real.Sup_le (set.forall_range_iff.2 hS) ha + +/-- +As `0` is the default value for `real.Sup` of the empty set, it suffices to show that `S` is +bounded above by `0` to show that `Sup S ≤ 0`. +-/ +lemma Sup_nonpos (S : set ℝ) (hS : ∀ x ∈ S, x ≤ (0:ℝ)) : Sup S ≤ 0 := +real.Sup_le hS le_rfl + /-- As `0` is the default value for `real.Inf` of the empty set, it suffices to show that `S` is bounded below by `0` to show that `0 ≤ Inf S`. From d2d14e720caf3d62d691bf49642b22b3272de57a Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 26 May 2023 06:31:58 +0000 Subject: [PATCH 1083/1291] chore(measure_theory/function/lp_space): add and reorder monotonicity results (#19092) In #19083, I generalize the normed_space results to work without a tight bound. This is much easier if I have the more general monotonicity results available first. The `snorm'_le_nnreal_smul_snorm'_of_ae_le_mul` and `snorm_ess_sup_le_nnreal_smul_snorm_ess_sup_of_ae_le_mul` lemmas are new, and the `snorm_le_nnreal_smul_snorm_of_ae_le_mul` lemma has been adjusted to not rely on `snorm_const_smul`. All the other lemmas have just been reordered. --- src/measure_theory/function/lp_space.lean | 139 +++++++++++++--------- 1 file changed, 83 insertions(+), 56 deletions(-) diff --git a/src/measure_theory/function/lp_space.lean b/src/measure_theory/function/lp_space.lean index 918ddb4fa3478..540edc013e8fb 100644 --- a/src/measure_theory/function/lp_space.lean +++ b/src/measure_theory/function/lp_space.lean @@ -1241,6 +1241,89 @@ end end has_measurable_add +section monotonicity + +lemma snorm'_le_nnreal_smul_snorm'_of_ae_le_mul {f : α → F} {g : α → G} {c : ℝ≥0} + (h : ∀ᵐ x ∂μ, ‖f x‖₊ ≤ c * ‖g x‖₊) {p : ℝ} (hp : 0 < p) : + snorm' f p μ ≤ c • snorm' g p μ := +begin + simp_rw [snorm'], + rw [←ennreal.rpow_le_rpow_iff hp, ennreal.smul_def, smul_eq_mul, + ennreal.mul_rpow_of_nonneg _ _ hp.le], + simp_rw [←ennreal.rpow_mul, one_div, inv_mul_cancel hp.ne.symm, ennreal.rpow_one, + ennreal.coe_rpow_of_nonneg _ hp.le, ←lintegral_const_mul' _ _ ennreal.coe_ne_top, + ←ennreal.coe_mul], + apply lintegral_mono_ae, + simp_rw [ennreal.coe_le_coe, ←nnreal.mul_rpow, nnreal.rpow_le_rpow_iff hp], + exact h, +end + +lemma snorm_ess_sup_le_nnreal_smul_snorm_ess_sup_of_ae_le_mul {f : α → F} {g : α → G} {c : ℝ≥0} + (h : ∀ᵐ x ∂μ, ‖f x‖₊ ≤ c * ‖g x‖₊) : + snorm_ess_sup f μ ≤ c • snorm_ess_sup g μ := +calc ess_sup (λ x, (‖f x‖₊: ℝ≥0∞)) μ + ≤ ess_sup (λ x, (↑(c * ‖g x‖₊) : ℝ≥0∞)) μ + : ess_sup_mono_ae $ h.mono $ λ x hx, ennreal.coe_le_coe.mpr hx +... = ess_sup (λ x, (c * ‖g x‖₊ : ℝ≥0∞)) μ : by simp_rw ennreal.coe_mul +... = c • ess_sup (λ x, (‖g x‖₊ : ℝ≥0∞)) μ : ennreal.ess_sup_const_mul + +lemma snorm_le_nnreal_smul_snorm_of_ae_le_mul {f : α → F} {g : α → G} {c : ℝ≥0} + (h : ∀ᵐ x ∂μ, ‖f x‖₊ ≤ c * ‖g x‖₊) (p : ℝ≥0∞) : + snorm f p μ ≤ c • snorm g p μ := +begin + by_cases h0 : p = 0, + { simp [h0], }, + by_cases h_top : p = ∞, + { rw h_top, + exact snorm_ess_sup_le_nnreal_smul_snorm_ess_sup_of_ae_le_mul h, }, + simp_rw snorm_eq_snorm' h0 h_top, + exact snorm'_le_nnreal_smul_snorm'_of_ae_le_mul h (ennreal.to_real_pos h0 h_top), +end + +-- TODO: add the whole family of lemmas? +private lemma le_mul_iff_eq_zero_of_nonneg_of_neg_of_nonneg {α} [linear_ordered_semiring α] + {a b c : α} (ha : 0 ≤ a) (hb : b < 0) (hc : 0 ≤ c) : a ≤ b * c ↔ a = 0 ∧ c = 0 := +begin + split, + { intro h, + exact ⟨(h.trans (mul_nonpos_of_nonpos_of_nonneg hb.le hc)).antisymm ha, + (nonpos_of_mul_nonneg_right (ha.trans h) hb).antisymm hc⟩ }, + { rintro ⟨rfl, rfl⟩, + rw mul_zero, } +end + +/-- When `c` is negative, `‖f x‖ ≤ c * ‖g x‖` is nonsense and forces both `f` and `g` to have an +`snorm` of `0`. -/ +lemma snorm_eq_zero_and_zero_of_ae_le_mul_neg {f : α → F} {g : α → G} {c : ℝ} + (h : ∀ᵐ x ∂μ, ‖f x‖ ≤ c * ‖g x‖) (hc : c < 0) (p : ℝ≥0∞) : + snorm f p μ = 0 ∧ snorm g p μ = 0 := +begin + simp_rw [le_mul_iff_eq_zero_of_nonneg_of_neg_of_nonneg (norm_nonneg _) hc (norm_nonneg _), + norm_eq_zero, eventually_and] at h, + change f =ᵐ[μ] 0 ∧ g =ᵐ[μ] 0 at h, + simp [snorm_congr_ae h.1, snorm_congr_ae h.2], +end + +lemma snorm_le_mul_snorm_of_ae_le_mul {f : α → F} {g : α → G} {c : ℝ} + (h : ∀ᵐ x ∂μ, ‖f x‖ ≤ c * ‖g x‖) (p : ℝ≥0∞) : + snorm f p μ ≤ (ennreal.of_real c) * snorm g p μ := +snorm_le_nnreal_smul_snorm_of_ae_le_mul + (h.mono $ λ x hx, hx.trans $ mul_le_mul_of_nonneg_right c.le_coe_to_nnreal (norm_nonneg _)) _ + +lemma mem_ℒp.of_nnnorm_le_mul {f : α → E} {g : α → F} {c : ℝ≥0} + (hg : mem_ℒp g p μ) (hf : ae_strongly_measurable f μ) (hfg : ∀ᵐ x ∂μ, ‖f x‖₊ ≤ c * ‖g x‖₊) : + mem_ℒp f p μ := +⟨hf, (snorm_le_nnreal_smul_snorm_of_ae_le_mul hfg p).trans_lt $ + ennreal.mul_lt_top ennreal.coe_ne_top hg.snorm_ne_top⟩ + +lemma mem_ℒp.of_le_mul {f : α → E} {g : α → F} {c : ℝ} + (hg : mem_ℒp g p μ) (hf : ae_strongly_measurable f μ) (hfg : ∀ᵐ x ∂μ, ‖f x‖ ≤ c * ‖g x‖) : + mem_ℒp f p μ := +⟨hf, (snorm_le_mul_snorm_of_ae_le_mul hfg p).trans_lt $ + ennreal.mul_lt_top ennreal.of_real_ne_top hg.snorm_ne_top⟩ + +end monotonicity + section normed_space variables {𝕜 : Type*} [normed_field 𝕜] [normed_space 𝕜 E] [normed_space 𝕜 F] @@ -1415,62 +1498,6 @@ by { apply hf.smul hφ, simp only [ennreal.div_top, add_zero] } end normed_space -section monotonicity - -lemma snorm_le_nnreal_smul_snorm_of_ae_le_mul {f : α → F} {g : α → G} {c : ℝ≥0} - (h : ∀ᵐ x ∂μ, ‖f x‖₊ ≤ c * ‖g x‖₊) (p : ℝ≥0∞) : - snorm f p μ ≤ c • snorm g p μ := -begin - rw [← snorm_norm g, ←c.nnnorm_eq, ennreal.smul_def, smul_eq_mul, ← snorm_const_smul (c : ℝ)], - swap, apply_instance, - refine snorm_mono_nnnorm_ae (h.mono $ λ x hx, hx.trans_eq _), - rw [pi.smul_apply, smul_eq_mul, nnnorm_mul, nnreal.nnnorm_eq, nnnorm_norm], -end - --- TODO: add the whole family of lemmas? -private lemma le_mul_iff_eq_zero_of_nonneg_of_neg_of_nonneg {α} [linear_ordered_semiring α] - {a b c : α} (ha : 0 ≤ a) (hb : b < 0) (hc : 0 ≤ c) : a ≤ b * c ↔ a = 0 ∧ c = 0 := -begin - split, - { intro h, - exact ⟨(h.trans (mul_nonpos_of_nonpos_of_nonneg hb.le hc)).antisymm ha, - (nonpos_of_mul_nonneg_right (ha.trans h) hb).antisymm hc⟩ }, - { rintro ⟨rfl, rfl⟩, - rw mul_zero, } -end - -/-- When `c` is negative, `‖f x‖ ≤ c * ‖g x‖` is nonsense and forces both `f` and `g` to have an -`snorm` of `0`. -/ -lemma snorm_eq_zero_and_zero_of_ae_le_mul_neg {f : α → F} {g : α → G} {c : ℝ} - (h : ∀ᵐ x ∂μ, ‖f x‖ ≤ c * ‖g x‖) (hc : c < 0) (p : ℝ≥0∞) : - snorm f p μ = 0 ∧ snorm g p μ = 0 := -begin - simp_rw [le_mul_iff_eq_zero_of_nonneg_of_neg_of_nonneg (norm_nonneg _) hc (norm_nonneg _), - norm_eq_zero, eventually_and] at h, - change f =ᵐ[μ] 0 ∧ g =ᵐ[μ] 0 at h, - simp [snorm_congr_ae h.1, snorm_congr_ae h.2], -end - -lemma snorm_le_mul_snorm_of_ae_le_mul {f : α → F} {g : α → G} {c : ℝ} - (h : ∀ᵐ x ∂μ, ‖f x‖ ≤ c * ‖g x‖) (p : ℝ≥0∞) : - snorm f p μ ≤ (ennreal.of_real c) * snorm g p μ := -snorm_le_nnreal_smul_snorm_of_ae_le_mul - (h.mono $ λ x hx, hx.trans $ mul_le_mul_of_nonneg_right c.le_coe_to_nnreal (norm_nonneg _)) _ - -lemma mem_ℒp.of_nnnorm_le_mul {f : α → E} {g : α → F} {c : ℝ≥0} - (hg : mem_ℒp g p μ) (hf : ae_strongly_measurable f μ) (hfg : ∀ᵐ x ∂μ, ‖f x‖₊ ≤ c * ‖g x‖₊) : - mem_ℒp f p μ := -⟨hf, (snorm_le_nnreal_smul_snorm_of_ae_le_mul hfg p).trans_lt $ - ennreal.mul_lt_top ennreal.coe_ne_top hg.snorm_ne_top⟩ - -lemma mem_ℒp.of_le_mul {f : α → E} {g : α → F} {c : ℝ} - (hg : mem_ℒp g p μ) (hf : ae_strongly_measurable f μ) (hfg : ∀ᵐ x ∂μ, ‖f x‖ ≤ c * ‖g x‖) : - mem_ℒp f p μ := -⟨hf, (snorm_le_mul_snorm_of_ae_le_mul hfg p).trans_lt $ - ennreal.mul_lt_top ennreal.of_real_ne_top hg.snorm_ne_top⟩ - -end monotonicity - lemma snorm_indicator_ge_of_bdd_below (hp : p ≠ 0) (hp' : p ≠ ∞) {f : α → F} (C : ℝ≥0) {s : set α} (hs : measurable_set s) (hf : ∀ᵐ x ∂μ, x ∈ s → C ≤ ‖s.indicator f x‖₊) : From c2258f7bf086b17eac0929d635403780c39e239f Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Fri, 26 May 2023 06:31:59 +0000 Subject: [PATCH 1084/1291] chore(topology/continuous_function/ideals): generalize type class requirements (#19093) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit For the continuous functional calculus in non-unital algebras the subtype `C(R, R)₀ := { f : C(R, R) // f 0 = 0 }` will be useful, which is exactly `continuous_map.ideal_of_set`. However, it will be necessary to let `R` be `ℂ`, `ℝ` and `ℝ≥0`, and for the latter we need these more general type class assumptions. --- src/topology/continuous_function/ideals.lean | 5 +++-- 1 file changed, 3 insertions(+), 2 deletions(-) diff --git a/src/topology/continuous_function/ideals.lean b/src/topology/continuous_function/ideals.lean index ece1cf99f5e39..c457752b83dbf 100644 --- a/src/topology/continuous_function/ideals.lean +++ b/src/topology/continuous_function/ideals.lean @@ -14,7 +14,7 @@ import topology.algebra.module.character_space /-! # Ideals of continuous functions -For a topological ring `R` and a topological space `X` there is a Galois connection between +For a topological semiring `R` and a topological space `X` there is a Galois connection between `ideal C(X, R)` and `set X` given by sending each `I : ideal C(X, R)` to `{x : X | ∀ f ∈ I, f x = 0}ᶜ` and mapping `s : set X` to the ideal with carrier `{f : C(X, R) | ∀ x ∈ sᶜ, f x = 0}`, and we call these maps `continuous_map.set_of_ideal` and @@ -77,7 +77,8 @@ open topological_space section topological_ring -variables {X R : Type*} [topological_space X] [ring R] [topological_space R] [topological_ring R] +variables {X R : Type*} [topological_space X] [semiring R] +variables [topological_space R] [topological_semiring R] variable (R) From 42e9a1fd3a99e10f82830349ba7f4f10e8961c2a Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Fri, 26 May 2023 06:32:00 +0000 Subject: [PATCH 1085/1291] feat(analysis/calculus/bump_function_inner): add `real.smooth_transition.proj_Icc` (#19097) Also add `real.smooth_transition.continuous_at`. From the sphere eversion project --- src/analysis/calculus/bump_function_inner.lean | 13 +++++++++++++ src/data/set/intervals/proj_Icc.lean | 14 ++++++++++++++ 2 files changed, 27 insertions(+) diff --git a/src/analysis/calculus/bump_function_inner.lean b/src/analysis/calculus/bump_function_inner.lean index 2c1173af84735..2afda5effcfb0 100644 --- a/src/analysis/calculus/bump_function_inner.lean +++ b/src/analysis/calculus/bump_function_inner.lean @@ -236,6 +236,16 @@ zero_of_nonpos le_rfl @[simp] protected lemma one : smooth_transition 1 = 1 := one_of_one_le le_rfl +/-- Since `real.smooth_transition` is constant on $(-∞, 0]$ and $[1, ∞)$, applying it to the +projection of `x : ℝ` to $[0, 1]$ gives the same result as applying it to `x`. -/ +@[simp] protected lemma proj_Icc : + smooth_transition (proj_Icc (0 : ℝ) 1 zero_le_one x) = smooth_transition x := +begin + refine congr_fun (Icc_extend_eq_self zero_le_one smooth_transition (λ x hx, _) (λ x hx, _)) x, + { rw [smooth_transition.zero, zero_of_nonpos hx.le] }, + { rw [smooth_transition.one, one_of_one_le hx.le] } +end + lemma le_one (x : ℝ) : smooth_transition x ≤ 1 := (div_le_one (pos_denom x)).2 $ le_add_of_nonneg_right (nonneg _) @@ -260,6 +270,9 @@ smooth_transition.cont_diff.cont_diff_at protected lemma continuous : continuous smooth_transition := (@smooth_transition.cont_diff 0).continuous +protected lemma continuous_at : continuous_at smooth_transition x := +smooth_transition.continuous.continuous_at + end smooth_transition end real diff --git a/src/data/set/intervals/proj_Icc.lean b/src/data/set/intervals/proj_Icc.lean index 871779f42b151..e6cb3cff298e2 100644 --- a/src/data/set/intervals/proj_Icc.lean +++ b/src/data/set/intervals/proj_Icc.lean @@ -114,6 +114,20 @@ congr_arg f $ proj_Icc_of_mem h hx Icc_extend h f x = f x := congr_arg f $ proj_Icc_coe h x +/-- If `f : α → β` is a constant both on $(-∞, a]$ and on $[b, +∞)$, then the extension of this +function from $[a, b]$ to the whole line is equal to the original function. -/ +lemma Icc_extend_eq_self (f : α → β) (ha : ∀ x < a, f x = f a) (hb : ∀ x, b < x → f x = f b) : + Icc_extend h (f ∘ coe) = f := +begin + ext x, + cases lt_or_le x a with hxa hax, + { simp [Icc_extend_of_le_left _ _ hxa.le, ha x hxa] }, + { cases le_or_lt x b with hxb hbx, + { lift x to Icc a b using ⟨hax, hxb⟩, + rw [Icc_extend_coe] }, + { simp [Icc_extend_of_right_le _ _ hbx.le, hb x hbx] } } +end + end set open set From 76f9c990d4b7c3dd26b87c4c4b51759e249d9e66 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Fri, 26 May 2023 08:33:09 +0000 Subject: [PATCH 1086/1291] feat(topology/subset_properties): add `sigma_compact_space` instances (#19100) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Add instances for * `α × β`, `α ⊕ β`, `ulift α`; * `Π i : ι, π i`, assuming `finite ι`; * `Σ i : ι, π i`, assuming `countable ι`. In each case, all input topological spaces are also assumed to be σ-compact. Motivated by the `prod` instance in the sphere eversion project. --- src/topology/constructions.lean | 4 +++ src/topology/subset_properties.lean | 47 +++++++++++++++++++++++++++++ 2 files changed, 51 insertions(+) diff --git a/src/topology/constructions.lean b/src/topology/constructions.lean index c417015bb6d40..f993d015b760f 100644 --- a/src/topology/constructions.lean +++ b/src/topology/constructions.lean @@ -1306,6 +1306,10 @@ lemma embedding_ulift_down [topological_space α] : embedding (ulift.down : ulift.{v u} α → α) := ⟨⟨rfl⟩, ulift.down_injective⟩ +lemma ulift.closed_embedding_down [topological_space α] : + closed_embedding (ulift.down : ulift.{v u} α → α) := +⟨embedding_ulift_down, by simp only [ulift.down_surjective.range_eq, is_closed_univ]⟩ + instance [topological_space α] [discrete_topology α] : discrete_topology (ulift α) := embedding_ulift_down.discrete_topology diff --git a/src/topology/subset_properties.lean b/src/topology/subset_properties.lean index d37fdbf1062bc..b7da3d32318db 100644 --- a/src/topology/subset_properties.lean +++ b/src/topology/subset_properties.lean @@ -1295,6 +1295,53 @@ variable {α} lemma exists_mem_compact_covering (x : α) : ∃ n, x ∈ compact_covering α n := Union_eq_univ_iff.mp (Union_compact_covering α) x +instance [sigma_compact_space β] : sigma_compact_space (α × β) := +⟨⟨λ n, compact_covering α n ×ˢ compact_covering β n, + λ _, (is_compact_compact_covering _ _).prod (is_compact_compact_covering _ _), + by simp only [Union_prod_of_monotone (compact_covering_subset α) (compact_covering_subset β), + Union_compact_covering, univ_prod_univ]⟩⟩ + +instance [finite ι] [Π i, topological_space (π i)] [Π i, sigma_compact_space (π i)] : + sigma_compact_space (Π i, π i) := +begin + refine ⟨⟨λ n, set.pi univ (λ i, compact_covering (π i) n), + λ n, is_compact_univ_pi $ λ i, is_compact_compact_covering _ _, _⟩⟩, + rw [Union_univ_pi_of_monotone], + { simp only [Union_compact_covering, pi_univ] }, + { exact λ i, compact_covering_subset (π i) } +end + +instance [sigma_compact_space β] : sigma_compact_space (α ⊕ β) := +⟨⟨λ n, sum.inl '' compact_covering α n ∪ sum.inr '' compact_covering β n, + λ n, ((is_compact_compact_covering α n).image continuous_inl).union + ((is_compact_compact_covering β n).image continuous_inr), + by simp only [Union_union_distrib, ← image_Union, Union_compact_covering, image_univ, + range_inl_union_range_inr]⟩⟩ + +instance [countable ι] [Π i, topological_space (π i)] [Π i, sigma_compact_space (π i)] : + sigma_compact_space (Σ i, π i) := +begin + casesI is_empty_or_nonempty ι, + { apply_instance }, + { rcases exists_surjective_nat ι with ⟨f, hf⟩, + refine ⟨⟨λ n, ⋃ k ≤ n, sigma.mk (f k) '' compact_covering (π (f k)) n, λ n, _, _⟩⟩, + { refine (finite_le_nat _).is_compact_bUnion (λ k _, _), + exact (is_compact_compact_covering _ _).image continuous_sigma_mk }, + { simp only [Union_eq_univ_iff, sigma.forall, mem_Union, hf.forall], + intros k y, + rcases exists_mem_compact_covering y with ⟨n, hn⟩, + refine ⟨max k n, k, le_max_left _ _, mem_image_of_mem _ _⟩, + exact compact_covering_subset _ (le_max_right _ _) hn } } +end + +protected theorem closed_embedding.sigma_compact_space {e : β → α} (he : closed_embedding e) : + sigma_compact_space β := +⟨⟨λ n, e ⁻¹' compact_covering α n, λ n, he.is_compact_preimage (is_compact_compact_covering _ _), + by rw [← preimage_Union, Union_compact_covering, preimage_univ]⟩⟩ + +instance [sigma_compact_space β] : sigma_compact_space (ulift.{u} β) := + ulift.closed_embedding_down.sigma_compact_space + /-- If `α` is a `σ`-compact space, then a locally finite family of nonempty sets of `α` can have only countably many elements, `set.countable` version. -/ protected lemma locally_finite.countable_univ {ι : Type*} {f : ι → set α} (hf : locally_finite f) From a33d01f7f8d39883e1ab4b86c9bf21692f1d4356 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 26 May 2023 10:27:18 +0000 Subject: [PATCH 1087/1291] feat(analysis/inner_product_space/spectrum): add a decomposition instance (#14870) This `decomposition` instances inherits the noncomputability of `orthogonal_projection`, but it is at least defeq to something useful. --- .../inner_product_space/projection.lean | 65 +++++++++++++++++++ .../inner_product_space/spectrum.lean | 22 ++++++- 2 files changed, 86 insertions(+), 1 deletion(-) diff --git a/src/analysis/inner_product_space/projection.lean b/src/analysis/inner_product_space/projection.lean index 80d319de5fad2..42e7c5024ac76 100644 --- a/src/analysis/inner_product_space/projection.lean +++ b/src/analysis/inner_product_space/projection.lean @@ -3,6 +3,7 @@ Copyright (c) 2019 Zhouhang Zhou. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Zhouhang Zhou, Frédéric Dupuis, Heather Macbeth -/ +import algebra.direct_sum.decomposition import analysis.convex.basic import analysis.inner_product_space.orthogonal import analysis.inner_product_space.symmetric @@ -1221,6 +1222,70 @@ begin (complete_space_coe_iff_is_complete.mp infer_instance) end +open_locale direct_sum + +/-- If `x` lies within an orthogonal family `v`, it can be expressed as a sum of projections. -/ +lemma orthogonal_family.sum_projection_of_mem_supr [fintype ι] + {V : ι → submodule 𝕜 E} [∀ i, complete_space ↥(V i)] + (hV : orthogonal_family 𝕜 (λ i, V i) (λ i, (V i).subtypeₗᵢ)) (x : E) (hx : x ∈ supr V) : + ∑ i, (orthogonal_projection (V i) x : E) = x := +begin + refine submodule.supr_induction _ hx (λ i x hx, _) _ (λ x y hx hy, _), + { refine (finset.sum_eq_single_of_mem i (finset.mem_univ _) $ λ j _ hij, _).trans + (orthogonal_projection_eq_self_iff.mpr hx), + rw [orthogonal_projection_mem_subspace_orthogonal_complement_eq_zero, submodule.coe_zero], + exact hV.is_ortho hij.symm hx }, + { simp_rw [map_zero, submodule.coe_zero, finset.sum_const_zero] }, + { simp_rw [map_add, submodule.coe_add, finset.sum_add_distrib], + exact congr_arg2 (+) hx hy }, +end + +/-- If a family of submodules is orthogonal, then the `orthogonal_projection` on a direct sum +is just the coefficient of that direct sum. -/ +lemma orthogonal_family.projection_direct_sum_coe_add_hom [decidable_eq ι] + {V : ι → submodule 𝕜 E} (hV : orthogonal_family 𝕜 (λ i, V i) (λ i, (V i).subtypeₗᵢ)) + (x : ⨁ i, V i) (i : ι) [complete_space ↥(V i)] : + orthogonal_projection (V i) (direct_sum.coe_add_monoid_hom V x) = x i := +begin + induction x using direct_sum.induction_on with j x x y hx hy, + { simp }, + { simp_rw [direct_sum.coe_add_monoid_hom_of, direct_sum.of, dfinsupp.single_add_hom_apply], + obtain rfl | hij := decidable.eq_or_ne i j, + { rw [orthogonal_projection_mem_subspace_eq_self, dfinsupp.single_eq_same] }, + { rw [orthogonal_projection_mem_subspace_orthogonal_complement_eq_zero, + dfinsupp.single_eq_of_ne hij.symm], + exact hV.is_ortho hij.symm x.prop } }, + { simp_rw [map_add, dfinsupp.add_apply], + exact congr_arg2 (+) hx hy }, +end + +/-- If a family of submodules is orthogonal and they span the whole space, then the orthogonal +projection provides a means to decompose the space into its submodules. + +The projection function is `decompose V x i = orthogonal_projection (V i) x`. + +See note [reducible non-instances]. -/ +@[reducible] +def orthogonal_family.decomposition [decidable_eq ι] [fintype ι] {V : ι → submodule 𝕜 E} + [∀ i, complete_space ↥(V i)] + (hV : orthogonal_family 𝕜 (λ i, V i) (λ i, (V i).subtypeₗᵢ)) (h : supr V = ⊤) : + direct_sum.decomposition V := +{ decompose' := λ x, dfinsupp.equiv_fun_on_fintype.symm $ λ i, orthogonal_projection (V i) x, + left_inv := λ x, begin + dsimp only, + letI := λ i, classical.dec_eq (V i), + rw [direct_sum.coe_add_monoid_hom, direct_sum.to_add_monoid, dfinsupp.lift_add_hom_apply, + dfinsupp.sum_add_hom_apply, dfinsupp.sum_eq_sum_fintype], + { simp_rw [equiv.apply_symm_apply, add_submonoid_class.coe_subtype], + exact hV.sum_projection_of_mem_supr _ ((h.ge : _) submodule.mem_top),}, + { intro i, + exact map_zero _ }, + end, + right_inv := λ x, begin + dsimp only, + simp_rw [hV.projection_direct_sum_coe_add_hom, dfinsupp.equiv_fun_on_fintype_symm_coe], + end } + end orthogonal_family section orthonormal_basis diff --git a/src/analysis/inner_product_space/spectrum.lean b/src/analysis/inner_product_space/spectrum.lean index 525c6e21e8b0f..5d13cff964c7e 100644 --- a/src/analysis/inner_product_space/spectrum.lean +++ b/src/analysis/inner_product_space/spectrum.lean @@ -5,6 +5,7 @@ Authors: Heather Macbeth -/ import analysis.inner_product_space.rayleigh import analysis.inner_product_space.pi_L2 +import algebra.direct_sum.decomposition /-! # Spectral theory of self-adjoint operators @@ -130,8 +131,27 @@ show (⨆ μ : {μ // (eigenspace T μ) ≠ ⊥}, eigenspace T μ)ᗮ = ⊥, by rw [supr_ne_bot_subtype, hT.orthogonal_supr_eigenspaces_eq_bot] include dec_𝕜 +omit hT +/-- The eigenspaces of a self-adjoint operator on a finite-dimensional inner product space `E` gives +an internal direct sum decomposition of `E`. -/-- The eigenspaces of a self-adjoint operator on a finite-dimensional inner product space `E` give +Note this takes `hT` as a `fact` to allow it to be an instance. -/ +noncomputable instance direct_sum_decomposition [hT : fact T.is_symmetric] : + direct_sum.decomposition (λ μ : eigenvalues T, eigenspace T μ) := +begin + haveI h : ∀ μ : eigenvalues T, complete_space (eigenspace T μ) := λ μ, by apply_instance, + exact hT.out.orthogonal_family_eigenspaces'.decomposition + (submodule.orthogonal_eq_bot_iff.mp hT.out.orthogonal_supr_eigenspaces_eq_bot'), +end + +lemma direct_sum_decompose_apply [hT : fact T.is_symmetric] (x : E) (μ : eigenvalues T) : + direct_sum.decompose (λ μ : eigenvalues T, eigenspace T μ) x μ + = orthogonal_projection (eigenspace T μ) x := +rfl + +include hT + +/-- The eigenspaces of a self-adjoint operator on a finite-dimensional inner product space `E` gives an internal direct sum decomposition of `E`. -/ lemma direct_sum_is_internal : direct_sum.is_internal (λ μ : eigenvalues T, eigenspace T μ) := From 5bb9fffd23f9f65b367f5d451da18cc60bf47335 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Fri, 26 May 2023 13:18:47 +0000 Subject: [PATCH 1088/1291] feat(data/set/finite): add a version of `set.finite.bUnion` (#19098) Add `set.finite.Union` and `equiv.set_finite_iff`. From the sphere eversion project. --- src/data/set/finite.lean | 19 +++++++++++++++++++ 1 file changed, 19 insertions(+) diff --git a/src/data/set/finite.lean b/src/data/set/finite.lean index b7eb2f977fc42..d9918a209a506 100644 --- a/src/data/set/finite.lean +++ b/src/data/set/finite.lean @@ -393,6 +393,10 @@ end fintype_instances end set +lemma equiv.set_finite_iff {s : set α} {t : set β} (hst : s ≃ t) : + s.finite ↔ t.finite := +by simp_rw [← set.finite_coe_iff, hst.finite_iff] + /-! ### Finset -/ namespace finset @@ -594,6 +598,21 @@ theorem finite.sInter {α : Type*} {s : set (set α)} {t : set α} (ht : t ∈ s (hf : t.finite) : (⋂₀ s).finite := hf.subset (sInter_subset_of_mem ht) +/-- If sets `s i` are finite for all `i` from a finite set `t` and are empty for `i ∉ t`, then the +union `⋃ i, s i` is a finite set. -/ +lemma finite.Union {ι : Type*} {s : ι → set α} {t : set ι} (ht : t.finite) + (hs : ∀ i ∈ t, (s i).finite) (he : ∀ i ∉ t, s i = ∅) : + (⋃ i, s i).finite := +begin + suffices : (⋃ i, s i) ⊆ (⋃ i ∈ t, s i), + { exact (ht.bUnion hs).subset this, }, + refine Union_subset (λ i x hx, _), + by_cases hi : i ∈ t, + { exact mem_bUnion hi hx }, + { rw [he i hi, mem_empty_iff_false] at hx, + contradiction, }, +end + theorem finite.bind {α β} {s : set α} {f : α → set β} (h : s.finite) (hf : ∀ a ∈ s, (f a).finite) : (s >>= f).finite := h.bUnion hf From e0736bb5b48bdadbca19dbd857e12bee38ccfbb8 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 26 May 2023 16:13:43 +0000 Subject: [PATCH 1089/1291] refactor(measure_theory/function/lp_space): generalize actions from `normed_field` to `normed_ring` (#19083) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The motivation is that this makes it easier to work with integrals in non-commutative rings. This makes the proof of Hölder's inequality slightly more painful, as we can no longer use `simp_rw [norm_smul]` and have to make monotonicity arguments instead. `rel_congr` may be able to clean this up in mathlib3. The results in the `normed_space` section (including Hölder's inequality) have been largely moved to the `monotonicity` section, where they hold more generally for arbitrary binary functions. The results about scalar actions now follow as trivial special cases. This also makes the `fails_quickly` linter reject the `complete_space (Lp_meas F 𝕜 m p μ)` instance. Since Lean4 is around the corner and there are better debugging tools there, I think it's ok to just no-lint it. --- .../conditional_expectation/basic.lean | 7 +- .../function/continuous_map_dense.lean | 2 +- src/measure_theory/function/lp_space.lean | 326 +++++++++++------- 3 files changed, 215 insertions(+), 120 deletions(-) diff --git a/src/measure_theory/function/conditional_expectation/basic.lean b/src/measure_theory/function/conditional_expectation/basic.lean index 0cacb74285b49..6921a058ecaf7 100644 --- a/src/measure_theory/function/conditional_expectation/basic.lean +++ b/src/measure_theory/function/conditional_expectation/basic.lean @@ -506,6 +506,10 @@ instance [hm : fact (m ≤ m0)] [complete_space F] [hp : fact (1 ≤ p)] : complete_space (Lp_meas_subgroup F m p μ) := by { rw (Lp_meas_subgroup_to_Lp_trim_iso F p μ hm.elim).complete_space_iff, apply_instance, } +-- For now just no-lint this; lean4's tree-based logging will make this easier to debug. +-- One possible change might be to generalize `𝕜` from `is_R_or_C` to `normed_field`, as this +-- result may well hold there. +@[nolint fails_quickly] instance [hm : fact (m ≤ m0)] [complete_space F] [hp : fact (1 ≤ p)] : complete_space (Lp_meas F 𝕜 m p μ) := by { rw (Lp_meas_subgroup_to_Lp_meas_iso F 𝕜 p μ).symm.complete_space_iff, apply_instance, } @@ -1263,8 +1267,7 @@ lemma condexp_ind_smul_smul' [normed_space ℝ F] [smul_comm_class ℝ 𝕜 F] ( (hμs : μ s ≠ ∞) (c : 𝕜) (x : F) : condexp_ind_smul hm hs hμs (c • x) = c • condexp_ind_smul hm hs hμs x := by rw [condexp_ind_smul, condexp_ind_smul, to_span_singleton_smul', - (to_span_singleton ℝ x).smul_comp_LpL_apply c - ↑(condexp_L2 ℝ hm (indicator_const_Lp 2 hs hμs (1 : ℝ)))] + (to_span_singleton ℝ x).smul_comp_LpL c, smul_apply] lemma condexp_ind_smul_ae_eq_smul (hm : m ≤ m0) (hs : measurable_set s) (hμs : μ s ≠ ∞) (x : G) : condexp_ind_smul hm hs hμs x diff --git a/src/measure_theory/function/continuous_map_dense.lean b/src/measure_theory/function/continuous_map_dense.lean index f0d835f76c076..ab07dc7c853a4 100644 --- a/src/measure_theory/function/continuous_map_dense.lean +++ b/src/measure_theory/function/continuous_map_dense.lean @@ -109,7 +109,7 @@ begin { refine function.support_subset_iff'.2 (λ x hx, _), simp only [hgv hx, pi.zero_apply, zero_smul] }, have gc_mem : mem_ℒp (λ x, g x • c) p μ, - { apply mem_ℒp.smul_of_top_left (mem_ℒp_top_const _), + { refine mem_ℒp.smul_of_top_left (mem_ℒp_top_const _) _, refine ⟨g.continuous.ae_strongly_measurable, _⟩, have : snorm (v.indicator (λ x, (1 : ℝ))) p μ < ⊤, { refine (snorm_indicator_const_le _ _).trans_lt _, diff --git a/src/measure_theory/function/lp_space.lean b/src/measure_theory/function/lp_space.lean index 540edc013e8fb..3617d299a125a 100644 --- a/src/measure_theory/function/lp_space.lean +++ b/src/measure_theory/function/lp_space.lean @@ -391,6 +391,10 @@ lemma snorm_ess_sup_congr_ae {f g : α → F} (hfg : f =ᵐ[μ] g) : snorm_ess_sup f μ = snorm_ess_sup g μ := ess_sup_congr_ae (hfg.fun_comp (coe ∘ nnnorm)) +lemma snorm_ess_sup_mono_nnnorm_ae {f g : α → F} (hfg : ∀ᵐ x ∂μ, ‖f x‖₊ ≤ ‖g x‖₊) : + snorm_ess_sup f μ ≤ snorm_ess_sup g μ := +ess_sup_mono_ae $ hfg.mono $ λ x hx, ennreal.coe_le_coe.mpr hx + lemma snorm_mono_nnnorm_ae {f : α → F} {g : α → G} (h : ∀ᵐ x ∂μ, ‖f x‖₊ ≤ ‖g x‖₊) : snorm f p μ ≤ snorm g p μ := begin @@ -1322,95 +1326,59 @@ lemma mem_ℒp.of_le_mul {f : α → E} {g : α → F} {c : ℝ} ⟨hf, (snorm_le_mul_snorm_of_ae_le_mul hfg p).trans_lt $ ennreal.mul_lt_top ennreal.of_real_ne_top hg.snorm_ne_top⟩ -end monotonicity - -section normed_space - -variables {𝕜 : Type*} [normed_field 𝕜] [normed_space 𝕜 E] [normed_space 𝕜 F] - -lemma snorm'_const_smul {f : α → F} (c : 𝕜) (hq_pos : 0 < q) : - snorm' (c • f) q μ = (‖c‖₊ : ℝ≥0∞) * snorm' f q μ := -begin - rw snorm', - simp_rw [pi.smul_apply, nnnorm_smul, ennreal.coe_mul, - ennreal.mul_rpow_of_nonneg _ _ hq_pos.le], - suffices h_integral : ∫⁻ a, ↑(‖c‖₊) ^ q * ↑‖f a‖₊ ^ q ∂μ - = (‖c‖₊ : ℝ≥0∞)^q * ∫⁻ a, ‖f a‖₊ ^ q ∂μ, - { apply_fun (λ x, x ^ (1/q)) at h_integral, - rw [h_integral, ennreal.mul_rpow_of_nonneg _ _ (by simp [hq_pos.le] : 0 ≤ 1 / q)], - congr, - simp_rw [←ennreal.rpow_mul, one_div, mul_inv_cancel hq_pos.ne.symm, ennreal.rpow_one], }, - rw lintegral_const_mul', - rw ennreal.coe_rpow_of_nonneg _ hq_pos.le, - exact ennreal.coe_ne_top, -end - -lemma snorm_ess_sup_const_smul {f : α → F} (c : 𝕜) : - snorm_ess_sup (c • f) μ = (‖c‖₊ : ℝ≥0∞) * snorm_ess_sup f μ := -by simp_rw [snorm_ess_sup, pi.smul_apply, nnnorm_smul, ennreal.coe_mul, ennreal.ess_sup_const_mul] - -lemma snorm_const_smul {f : α → F} (c : 𝕜) : - snorm (c • f) p μ = (‖c‖₊ : ℝ≥0∞) * snorm f p μ := -begin - by_cases h0 : p = 0, - { simp [h0], }, - by_cases h_top : p = ∞, - { simp [h_top, snorm_ess_sup_const_smul], }, - repeat { rw snorm_eq_snorm' h0 h_top, }, - rw ←ne.def at h0, - exact snorm'_const_smul c (ennreal.to_real_pos h0 h_top), -end - -lemma mem_ℒp.const_smul {f : α → E} (hf : mem_ℒp f p μ) (c : 𝕜) : - mem_ℒp (c • f) p μ := -⟨ae_strongly_measurable.const_smul hf.1 c, - (snorm_const_smul c).le.trans_lt (ennreal.mul_lt_top ennreal.coe_ne_top hf.2.ne)⟩ - -lemma mem_ℒp.const_mul {f : α → 𝕜} (hf : mem_ℒp f p μ) (c : 𝕜) : - mem_ℒp (λ x, c * f x) p μ := -hf.const_smul c - -lemma snorm'_smul_le_mul_snorm' {p q r : ℝ} - {f : α → E} (hf : ae_strongly_measurable f μ) {φ : α → 𝕜} (hφ : ae_strongly_measurable φ μ) +lemma snorm'_le_snorm'_mul_snorm' {p q r : ℝ} + {f : α → E} (hf : ae_strongly_measurable f μ) {g : α → F} (hg : ae_strongly_measurable g μ) + (b : E → F → G) (h : ∀ᵐ x ∂μ, ‖b (f x) (g x)‖₊ ≤ ‖f x‖₊ * ‖g x‖₊) (hp0_lt : 0 < p) (hpq : p < q) (hpqr : 1/p = 1/q + 1/r) : - snorm' (φ • f) p μ ≤ snorm' φ q μ * snorm' f r μ := + snorm' (λ x, b (f x) (g x)) p μ ≤ snorm' f q μ * snorm' g r μ := begin - simp_rw [snorm', pi.smul_apply', nnnorm_smul, ennreal.coe_mul], - exact ennreal.lintegral_Lp_mul_le_Lq_mul_Lr hp0_lt hpq hpqr μ hφ.ennnorm - hf.ennnorm, -end - -lemma snorm_smul_le_snorm_top_mul_snorm (p : ℝ≥0∞) - {f : α → E} (hf : ae_strongly_measurable f μ) (φ : α → 𝕜) : - snorm (φ • f) p μ ≤ snorm φ ∞ μ * snorm f p μ := + rw snorm', + calc (∫⁻ (a : α), ↑‖b (f a) (g a)‖₊ ^ p ∂μ) ^ (1 / p) + ≤ (∫⁻ (a : α), ↑(‖f a‖₊ * ‖g a‖₊) ^ p ∂μ) ^ (1 / p) : + (ennreal.rpow_le_rpow_iff $ one_div_pos.mpr (hp0_lt)).mpr $ + lintegral_mono_ae $ h.mono $ λ a ha, (ennreal.rpow_le_rpow_iff (hp0_lt)).mpr $ + ennreal.coe_le_coe.mpr $ ha + ... ≤ _ : _, + simp_rw [snorm', ennreal.coe_mul], + exact ennreal.lintegral_Lp_mul_le_Lq_mul_Lr hp0_lt hpq hpqr μ hf.ennnorm + hg.ennnorm, +end + +lemma snorm_le_snorm_top_mul_snorm (p : ℝ≥0∞) + (f : α → E) {g : α → F} (hg : ae_strongly_measurable g μ) (b : E → F → G) + (h : ∀ᵐ x ∂μ, ‖b (f x) (g x)‖₊ ≤ ‖f x‖₊ * ‖g x‖₊) : + snorm (λ x, b (f x) (g x)) p μ ≤ snorm f ∞ μ * snorm g p μ := begin by_cases hp_top : p = ∞, - { simp_rw [hp_top, snorm_exponent_top, snorm_ess_sup, pi.smul_apply', nnnorm_smul, - ennreal.coe_mul], - exact ennreal.ess_sup_mul_le _ _, }, + { simp_rw [hp_top, snorm_exponent_top], + refine le_trans (ess_sup_mono_ae $ h.mono $ λ a ha, _) (ennreal.ess_sup_mul_le _ _), + simp_rw [pi.mul_apply, ←ennreal.coe_mul, ennreal.coe_le_coe], + exact ha }, by_cases hp_zero : p = 0, { simp only [hp_zero, snorm_exponent_zero, mul_zero, le_zero_iff], }, simp_rw [snorm_eq_lintegral_rpow_nnnorm hp_zero hp_top, snorm_exponent_top, snorm_ess_sup], - calc (∫⁻ x, ↑‖(φ • f) x‖₊ ^ p.to_real ∂μ) ^ (1 / p.to_real) - = (∫⁻ x, ↑‖φ x‖₊ ^ p.to_real * ↑‖f x‖₊ ^ p.to_real ∂μ) ^ (1 / p.to_real) : + calc (∫⁻ x, ↑‖b (f x) (g x)‖₊ ^ p.to_real ∂μ) ^ (1 / p.to_real) + ≤ (∫⁻ x, ↑‖f x‖₊ ^ p.to_real * ↑‖g x‖₊ ^ p.to_real ∂μ) ^ (1 / p.to_real) : begin - congr, - ext1 x, - rw [pi.smul_apply', nnnorm_smul, ennreal.coe_mul, - ennreal.mul_rpow_of_nonneg _ _ (ennreal.to_real_nonneg)], + refine ennreal.rpow_le_rpow _ (one_div_nonneg.mpr ennreal.to_real_nonneg), + refine lintegral_mono_ae (h.mono $ λ a ha, _), + rw ←ennreal.mul_rpow_of_nonneg _ _ ennreal.to_real_nonneg, + refine ennreal.rpow_le_rpow _ ennreal.to_real_nonneg, + rw [←ennreal.coe_mul, ennreal.coe_le_coe], + exact ha, end - ... ≤ (∫⁻ x, (ess_sup (λ x, ↑‖φ x‖₊) μ) ^ p.to_real * ↑‖f x‖₊ ^ p.to_real ∂μ) ^ (1 / p.to_real) : + ... ≤ (∫⁻ x, (ess_sup (λ x, ↑‖f x‖₊) μ) ^ p.to_real * ↑‖g x‖₊ ^ p.to_real ∂μ) ^ (1 / p.to_real) : begin refine ennreal.rpow_le_rpow _ _, swap, { rw one_div_nonneg, exact ennreal.to_real_nonneg, }, refine lintegral_mono_ae _, - filter_upwards [@ennreal.ae_le_ess_sup _ _ μ (λ x, ↑‖φ x‖₊)] with x hx, + filter_upwards [@ennreal.ae_le_ess_sup _ _ μ (λ x, ↑‖f x‖₊)] with x hx, exact mul_le_mul_right' (ennreal.rpow_le_rpow hx ennreal.to_real_nonneg) _ end - ... = ess_sup (λ x, ↑‖φ x‖₊) μ * (∫⁻ x, ↑‖f x‖₊ ^ p.to_real ∂μ) ^ (1 / p.to_real) : + ... = ess_sup (λ x, ↑‖f x‖₊) μ * (∫⁻ x, ↑‖g x‖₊ ^ p.to_real ∂μ) ^ (1 / p.to_real) : begin rw lintegral_const_mul'', - swap, { exact hf.nnnorm.ae_measurable.coe_nnreal_ennreal.pow ae_measurable_const, }, + swap, { exact hg.nnnorm.ae_measurable.coe_nnreal_ennreal.pow ae_measurable_const, }, rw ennreal.mul_rpow_of_nonneg, swap, { rw one_div_nonneg, exact ennreal.to_real_nonneg, }, rw [← ennreal.rpow_mul, one_div, mul_inv_cancel, ennreal.rpow_one], @@ -1419,26 +1387,26 @@ begin end end -lemma snorm_smul_le_snorm_mul_snorm_top (p : ℝ≥0∞) - (f : α → E) {φ : α → 𝕜} (hφ : ae_strongly_measurable φ μ) : - snorm (φ • f) p μ ≤ snorm φ p μ * snorm f ∞ μ := +lemma snorm_le_snorm_mul_snorm_top (p : ℝ≥0∞) + {f : α → E} (hf : ae_strongly_measurable f μ) (g : α → F) (b : E → F → G) + (h : ∀ᵐ x ∂μ, ‖b (f x) (g x)‖₊ ≤ ‖f x‖₊ * ‖g x‖₊) : + snorm (λ x, b (f x) (g x)) p μ ≤ snorm f p μ * snorm g ∞ μ := begin - rw ← snorm_norm, - simp_rw [pi.smul_apply', norm_smul], - have : (λ x, ‖φ x‖ * ‖f x‖) = (λ x, ‖f x‖) • (λ x, ‖φ x‖), - { rw [smul_eq_mul, mul_comm], refl, }, - rw this, - have h := snorm_smul_le_snorm_top_mul_snorm p hφ.norm (λ x, ‖f x‖), - refine h.trans_eq _, - simp_rw snorm_norm, + rw [← snorm_norm f, ← snorm_norm g], + refine (snorm_mono_ae_real h).trans _, + simp_rw [mul_comm ‖f _‖₊, nnreal.coe_mul, coe_nnnorm], rw mul_comm, + refine snorm_le_snorm_top_mul_snorm p (λ x, ‖g x‖) hf.norm _ (h.mono $ λ x hx, _), + simp_rw [nnnorm_mul], end -/-- Hölder's inequality, as an inequality on the `ℒp` seminorm of a scalar product `φ • f`. -/ -lemma snorm_smul_le_mul_snorm {p q r : ℝ≥0∞} - {f : α → E} (hf : ae_strongly_measurable f μ) {φ : α → 𝕜} (hφ : ae_strongly_measurable φ μ) +/-- Hölder's inequality, as an inequality on the `ℒp` seminorm of an elementwise operation +`λ x, b (f x) (g x)`. -/ +lemma snorm_le_snorm_mul_snorm_of_nnnorm {p q r : ℝ≥0∞} + {f : α → E} (hf : ae_strongly_measurable f μ) {g : α → F} (hg : ae_strongly_measurable g μ) + (b : E → F → G) (h : ∀ᵐ x ∂μ, ‖b (f x) (g x)‖₊ ≤ ‖f x‖₊ * ‖g x‖₊) (hpqr : 1/p = 1/q + 1/r) : - snorm (φ • f) p μ ≤ snorm φ q μ * snorm f r μ := + snorm (λ x, b (f x) (g x)) p μ ≤ snorm f q μ * snorm g r μ := begin by_cases hp_zero : p = 0, { simp [hp_zero], }, @@ -1456,12 +1424,12 @@ begin { have hpr : p = r, { simpa only [hq_top, one_div, ennreal.div_top, zero_add, inv_inj] using hpqr, }, rw [← hpr, hq_top], - exact snorm_smul_le_snorm_top_mul_snorm p hf φ, }, + exact snorm_le_snorm_top_mul_snorm p f hg b h, }, by_cases hr_top : r = ∞, { have hpq : p = q, { simpa only [hr_top, one_div, ennreal.div_top, add_zero, inv_inj] using hpqr, }, rw [← hpq, hr_top], - exact snorm_smul_le_snorm_mul_snorm_top p f hφ, }, + exact snorm_le_snorm_mul_snorm_top p hf g b h, }, have hpq : p < q, { suffices : 1 / q < 1 / p, { rwa [one_div, one_div, ennreal.inv_lt_inv] at this, }, @@ -1471,7 +1439,7 @@ begin { simp only [hr_top, one_div, ne.def, ennreal.inv_eq_zero, not_false_iff], }, }, rw [snorm_eq_snorm' hp_zero (hpq.trans_le le_top).ne, snorm_eq_snorm' hq_ne_zero hq_top, snorm_eq_snorm' hr_ne_zero hr_top], - refine snorm'_smul_le_mul_snorm' hf hφ _ _ _, + refine snorm'_le_snorm'_mul_snorm' hf hg _ h _ _ _, { exact ennreal.to_real_pos hp_zero (hpq.trans_le le_top).ne, }, { exact ennreal.to_real_strict_mono hq_top hpq, }, rw [← ennreal.one_to_real, ← ennreal.to_real_div, ← ennreal.to_real_div, ← ennreal.to_real_div, @@ -1480,6 +1448,74 @@ begin { simp only [hr_ne_zero, one_div, ne.def, ennreal.inv_eq_top, not_false_iff], }, end +/-- Hölder's inequality, as an inequality on the `ℒp` seminorm of an elementwise operation +`λ x, b (f x) (g x)`. -/ +lemma snorm_le_snorm_mul_snorm'_of_norm {p q r : ℝ≥0∞} + {f : α → E} (hf : ae_strongly_measurable f μ) {g : α → F} (hg : ae_strongly_measurable g μ) + (b : E → F → G) (h : ∀ᵐ x ∂μ, ‖b (f x) (g x)‖ ≤ ‖f x‖ * ‖g x‖) + (hpqr : 1/p = 1/q + 1/r) : + snorm (λ x, b (f x) (g x)) p μ ≤ snorm f q μ * snorm g r μ := +snorm_le_snorm_mul_snorm_of_nnnorm hf hg b h hpqr + +end monotonicity + +/-! +### Bounded actions by normed rings + +In this section we show inequalities on the norm. +-/ +section has_bounded_smul + +variables {𝕜 : Type*} [normed_ring 𝕜] [mul_action_with_zero 𝕜 E] [mul_action_with_zero 𝕜 F] +variables [has_bounded_smul 𝕜 E] [has_bounded_smul 𝕜 F] + +lemma snorm'_const_smul_le (c : 𝕜) (f : α → F) (hq_pos : 0 < q) : + snorm' (c • f) q μ ≤ ‖c‖₊ • snorm' f q μ := +snorm'_le_nnreal_smul_snorm'_of_ae_le_mul (eventually_of_forall $ λ a, nnnorm_smul_le _ _) hq_pos + +lemma snorm_ess_sup_const_smul_le (c : 𝕜) (f : α → F) : + snorm_ess_sup (c • f) μ ≤ ‖c‖₊ • snorm_ess_sup f μ := +snorm_ess_sup_le_nnreal_smul_snorm_ess_sup_of_ae_le_mul + (eventually_of_forall $ λ a, nnnorm_smul_le _ _) + +lemma snorm_const_smul_le (c : 𝕜) (f : α → F) : + snorm (c • f) p μ ≤ ‖c‖₊ • snorm f p μ := +snorm_le_nnreal_smul_snorm_of_ae_le_mul (eventually_of_forall $ λ a, nnnorm_smul_le _ _) _ + +lemma mem_ℒp.const_smul {f : α → E} (hf : mem_ℒp f p μ) (c : 𝕜) : + mem_ℒp (c • f) p μ := +⟨ae_strongly_measurable.const_smul hf.1 c, + (snorm_const_smul_le c f).trans_lt (ennreal.mul_lt_top ennreal.coe_ne_top hf.2.ne)⟩ + +lemma mem_ℒp.const_mul {R} [normed_ring R] {f : α → R} (hf : mem_ℒp f p μ) (c : R) : + mem_ℒp (λ x, c * f x) p μ := +hf.const_smul c + +lemma snorm'_smul_le_mul_snorm' {p q r : ℝ} + {f : α → E} (hf : ae_strongly_measurable f μ) {φ : α → 𝕜} (hφ : ae_strongly_measurable φ μ) + (hp0_lt : 0 < p) (hpq : p < q) (hpqr : 1/p = 1/q + 1/r) : + snorm' (φ • f) p μ ≤ snorm' φ q μ * snorm' f r μ := +snorm'_le_snorm'_mul_snorm' hφ hf (•) + (eventually_of_forall $ λ a, nnnorm_smul_le _ _) hp0_lt hpq hpqr + +lemma snorm_smul_le_snorm_top_mul_snorm (p : ℝ≥0∞) + {f : α → E} (hf : ae_strongly_measurable f μ) (φ : α → 𝕜) : + snorm (φ • f) p μ ≤ snorm φ ∞ μ * snorm f p μ := +(snorm_le_snorm_top_mul_snorm p φ hf (•) (eventually_of_forall $ λ a, nnnorm_smul_le _ _) : _) + +lemma snorm_smul_le_snorm_mul_snorm_top (p : ℝ≥0∞) + (f : α → E) {φ : α → 𝕜} (hφ : ae_strongly_measurable φ μ) : + snorm (φ • f) p μ ≤ snorm φ p μ * snorm f ∞ μ := +(snorm_le_snorm_mul_snorm_top p hφ f (•) (eventually_of_forall $ λ a, nnnorm_smul_le _ _) : _) + +/-- Hölder's inequality, as an inequality on the `ℒp` seminorm of a scalar product `φ • f`. -/ +lemma snorm_smul_le_mul_snorm {p q r : ℝ≥0∞} + {f : α → E} (hf : ae_strongly_measurable f μ) {φ : α → 𝕜} (hφ : ae_strongly_measurable φ μ) + (hpqr : 1/p = 1/q + 1/r) : + snorm (φ • f) p μ ≤ snorm φ q μ * snorm f r μ := +(snorm_le_snorm_mul_snorm_of_nnnorm hφ hf (•) + (eventually_of_forall $ λ a, nnnorm_smul_le _ _) hpqr : _) + lemma mem_ℒp.smul {p q r : ℝ≥0∞} {f : α → E} {φ : α → 𝕜} (hf : mem_ℒp f r μ) (hφ : mem_ℒp φ q μ) (hpqr : 1/p = 1/q + 1/r) : mem_ℒp (φ • f) p μ := @@ -1496,6 +1532,42 @@ lemma mem_ℒp.smul_of_top_left {p : ℝ≥0∞} {f : α → E} {φ : α → mem_ℒp (φ • f) p μ := by { apply hf.smul hφ, simp only [ennreal.div_top, add_zero] } +end has_bounded_smul + +/-! +### Bounded actions by normed division rings + +The inequalities in the previous section are now tight. +-/ +section normed_space + +variables {𝕜 : Type*} [normed_division_ring 𝕜] [mul_action_with_zero 𝕜 E] [module 𝕜 F] +variables [has_bounded_smul 𝕜 E] [has_bounded_smul 𝕜 F] + +lemma snorm'_const_smul {f : α → F} (c : 𝕜) (hq_pos : 0 < q) : + snorm' (c • f) q μ = ‖c‖₊ • snorm' f q μ := +begin + obtain rfl | hc := eq_or_ne c 0, + { simp [snorm', hq_pos], }, + refine le_antisymm (snorm'_const_smul_le _ _ hq_pos) _, + have : snorm' _ q μ ≤ _:= snorm'_const_smul_le (c⁻¹) (c • f) hq_pos, + rwa [inv_smul_smul₀ hc, nnnorm_inv, ennreal.le_inv_smul_iff (nnnorm_ne_zero_iff.mpr hc)] at this, +end + +lemma snorm_ess_sup_const_smul (c : 𝕜) (f : α → F) : + snorm_ess_sup (c • f) μ = (‖c‖₊ : ℝ≥0∞) * snorm_ess_sup f μ := +by simp_rw [snorm_ess_sup, pi.smul_apply, nnnorm_smul, ennreal.coe_mul, ennreal.ess_sup_const_mul] + +lemma snorm_const_smul (c : 𝕜) (f : α → F) : + snorm (c • f) p μ = (‖c‖₊ : ℝ≥0∞) * snorm f p μ := +begin + obtain rfl | hc := eq_or_ne c 0, + { simp, }, + refine le_antisymm (snorm_const_smul_le _ _) _, + have : snorm _ p μ ≤ _:= snorm_const_smul_le (c⁻¹) (c • f), + rwa [inv_smul_smul₀ hc, nnnorm_inv, ennreal.le_inv_smul_iff (nnnorm_ne_zero_iff.mpr hc)] at this, +end + end normed_space lemma snorm_indicator_ge_of_bdd_below (hp : p ≠ 0) (hp' : p ≠ ∞) @@ -1908,14 +1980,17 @@ example [fact (1 ≤ p)] : seminormed_add_group.to_has_nnnorm = (Lp.has_nnnorm : has_nnnorm (Lp E p μ)) := rfl -section normed_space +section has_bounded_smul -variables {𝕜 : Type*} [normed_field 𝕜] [normed_space 𝕜 E] +variables {𝕜 𝕜' : Type*} +variables [normed_ring 𝕜] [normed_ring 𝕜'] [module 𝕜 E] [module 𝕜' E] +variables [has_bounded_smul 𝕜 E] [has_bounded_smul 𝕜' E] lemma mem_Lp_const_smul (c : 𝕜) (f : Lp E p μ) : c • ↑f ∈ Lp E p μ := begin - rw [mem_Lp_iff_snorm_lt_top, snorm_congr_ae (ae_eq_fun.coe_fn_smul _ _), snorm_const_smul, - ennreal.mul_lt_top_iff], + rw [mem_Lp_iff_snorm_lt_top, snorm_congr_ae (ae_eq_fun.coe_fn_smul _ _)], + refine (snorm_const_smul_le _ _).trans_lt _, + rw [ennreal.smul_def, smul_eq_mul, ennreal.mul_lt_top_iff], exact or.inl ⟨ennreal.coe_lt_top, f.prop⟩, end @@ -1936,20 +2011,40 @@ instance : module 𝕜 (Lp E p μ) := lemma coe_fn_smul (c : 𝕜) (f : Lp E p μ) : ⇑(c • f) =ᵐ[μ] c • f := ae_eq_fun.coe_fn_smul _ _ -lemma norm_const_smul (c : 𝕜) (f : Lp E p μ) : ‖c • f‖ = ‖c‖ * ‖f‖ := -by rw [norm_def, snorm_congr_ae (coe_fn_smul _ _), snorm_const_smul c, - ennreal.to_real_mul, ennreal.coe_to_real, coe_nnnorm, norm_def] +instance [module 𝕜ᵐᵒᵖ E] [has_bounded_smul 𝕜ᵐᵒᵖ E] [is_central_scalar 𝕜 E] : + is_central_scalar 𝕜 (Lp E p μ) := +{ op_smul_eq_smul := λ k f, subtype.ext $ op_smul_eq_smul k (f : α →ₘ[μ] E) } + +instance [smul_comm_class 𝕜 𝕜' E] : smul_comm_class 𝕜 𝕜' (Lp E p μ) := +{ smul_comm := λ k k' f, subtype.ext $ smul_comm k k' (f : α →ₘ[μ] E) } + +instance [has_smul 𝕜 𝕜'] [is_scalar_tower 𝕜 𝕜' E] : is_scalar_tower 𝕜 𝕜' (Lp E p μ) := +{ smul_assoc := λ k k' f, subtype.ext $ smul_assoc k k' (f : α →ₘ[μ] E) } + +instance [fact (1 ≤ p)] : has_bounded_smul 𝕜 (Lp E p μ) := +-- TODO: add `has_bounded_smul.of_nnnorm_smul_le +has_bounded_smul.of_norm_smul_le $ λ r f, begin + suffices : (‖r • f‖₊ : ℝ≥0∞) ≤ ‖r‖₊ * ‖f‖₊, + { exact_mod_cast this }, + rw [nnnorm_def, nnnorm_def, ennreal.coe_to_nnreal (Lp.snorm_ne_top _), + snorm_congr_ae (coe_fn_smul _ _), ennreal.coe_to_nnreal (Lp.snorm_ne_top _)], + exact snorm_const_smul_le r f, +end + +end has_bounded_smul + +section normed_space +variables {𝕜 : Type*} [normed_field 𝕜] [normed_space 𝕜 E] instance [fact (1 ≤ p)] : normed_space 𝕜 (Lp E p μ) := -{ norm_smul_le := λ _ _, by simp [norm_const_smul] } +{ norm_smul_le := λ _ _, norm_smul_le _ _ } end normed_space end Lp namespace mem_ℒp - -variables {𝕜 : Type*} [normed_field 𝕜] [normed_space 𝕜 E] +variables {𝕜 : Type*} [normed_ring 𝕜] [module 𝕜 E] [has_bounded_smul 𝕜 E] lemma to_Lp_const_smul {f : α → E} (c : 𝕜) (hf : mem_ℒp f p μ) : (hf.const_smul c).to_Lp (c • f) = c • hf.to_Lp f := rfl @@ -2276,11 +2371,11 @@ lemma lipschitz_with.comp_mem_ℒp {α E F} {K} [measurable_space α] {μ : meas [normed_add_comm_group E] [normed_add_comm_group F] {f : α → E} {g : E → F} (hg : lipschitz_with K g) (g0 : g 0 = 0) (hL : mem_ℒp f p μ) : mem_ℒp (g ∘ f) p μ := begin - have : ∀ᵐ x ∂μ, ‖g (f x)‖ ≤ K * ‖f x‖, - { apply filter.eventually_of_forall (λ x, _), - rw [← dist_zero_right, ← dist_zero_right, ← g0], - apply hg.dist_le_mul }, - exact hL.of_le_mul (hg.continuous.comp_ae_strongly_measurable hL.1) this, + have : ∀ x, ‖g (f x)‖ ≤ K * ‖f x‖, + { intro a, + -- TODO: add `lipschitz_with.nnnorm_sub_le` and `lipschitz_with.nnnorm_le` + simpa [g0] using hg.norm_sub_le (f a) 0, }, + exact hL.of_le_mul (hg.continuous.comp_ae_strongly_measurable hL.1) (eventually_of_forall this), end lemma measure_theory.mem_ℒp.of_comp_antilipschitz_with {α E F} {K'} @@ -2288,13 +2383,14 @@ lemma measure_theory.mem_ℒp.of_comp_antilipschitz_with {α E F} {K'} {f : α → E} {g : E → F} (hL : mem_ℒp (g ∘ f) p μ) (hg : uniform_continuous g) (hg' : antilipschitz_with K' g) (g0 : g 0 = 0) : mem_ℒp f p μ := begin - have A : ∀ᵐ x ∂μ, ‖f x‖ ≤ K' * ‖g (f x)‖, - { apply filter.eventually_of_forall (λ x, _), + have A : ∀ x, ‖f x‖ ≤ K' * ‖g (f x)‖, + { intro x, + -- TODO: add `antilipschitz_with.le_mul_nnnorm_sub` and `antilipschitz_with.le_mul_norm` rw [← dist_zero_right, ← dist_zero_right, ← g0], apply hg'.le_mul_dist }, have B : ae_strongly_measurable f μ := ((hg'.uniform_embedding hg).embedding.ae_strongly_measurable_comp_iff.1 hL.1), - exact hL.of_le_mul B A, + exact hL.of_le_mul B (filter.eventually_of_forall A), end namespace lipschitz_with @@ -2410,7 +2506,8 @@ begin refl, end -lemma smul_comp_Lp {𝕜'} [normed_field 𝕜'] [normed_space 𝕜' F] [smul_comm_class 𝕜 𝕜' F] +lemma smul_comp_Lp {𝕜'} [normed_ring 𝕜'] [module 𝕜' F] [has_bounded_smul 𝕜' F] + [smul_comm_class 𝕜 𝕜' F] (c : 𝕜') (L : E →L[𝕜] F) (f : Lp E p μ) : (c • L).comp_Lp f = c • L.comp_Lp f := begin @@ -2467,16 +2564,11 @@ lemma add_comp_LpL [fact (1 ≤ p)] (L L' : E →L[𝕜] F) : (L + L').comp_LpL p μ = L.comp_LpL p μ + L'.comp_LpL p μ := by { ext1 f, exact add_comp_Lp L L' f } -lemma smul_comp_LpL [fact (1 ≤ p)] (c : 𝕜) (L : E →L[𝕜] F) : - (c • L).comp_LpL p μ = c • (L.comp_LpL p μ) := +lemma smul_comp_LpL [fact (1 ≤ p)] {𝕜'} [normed_ring 𝕜'] [module 𝕜' F] + [has_bounded_smul 𝕜' F] [smul_comm_class 𝕜 𝕜' F] (c : 𝕜') (L : E →L[𝕜] F) : + (c • L).comp_LpL p μ = c • L.comp_LpL p μ := by { ext1 f, exact smul_comp_Lp c L f } -/-- TODO: written in an "apply" way because of a missing `has_smul` instance. -/ -lemma smul_comp_LpL_apply [fact (1 ≤ p)] {𝕜'} [normed_field 𝕜'] [normed_space 𝕜' F] - [smul_comm_class 𝕜 𝕜' F] (c : 𝕜') (L : E →L[𝕜] F) (f : Lp E p μ) : - (c • L).comp_LpL p μ f = c • (L.comp_LpL p μ f) := -smul_comp_Lp c L f - lemma norm_compLpL_le [fact (1 ≤ p)] (L : E →L[𝕜] F) : ‖L.comp_LpL p μ‖ ≤ ‖L‖ := linear_map.mk_continuous_norm_le _ (norm_nonneg _) _ From 373b03b5b9d0486534edbe94747f23cb3712f93d Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Fri, 26 May 2023 16:13:44 +0000 Subject: [PATCH 1090/1291] feat(analysis/convex/gauge): gauge of a convex nhd of zero is continuous (#19102) From the Brouwer Fixed Point Theorem project. Co-authored-by: @Shamrock-Frost --- src/analysis/convex/gauge.lean | 39 ++++++++++++++++++---------------- 1 file changed, 21 insertions(+), 18 deletions(-) diff --git a/src/analysis/convex/gauge.lean b/src/analysis/convex/gauge.lean index e027faa1a0d3d..2304cd00b55fe 100644 --- a/src/analysis/convex/gauge.lean +++ b/src/analysis/convex/gauge.lean @@ -40,7 +40,7 @@ Minkowski functional, gauge -/ open normed_field set -open_locale pointwise +open_locale pointwise topology nnreal noncomputable theory @@ -438,23 +438,7 @@ section norm variables [seminormed_add_comm_group E] [normed_space ℝ E] {s : set E} {r : ℝ} {x : E} lemma gauge_unit_ball (x : E) : gauge (metric.ball (0 : E) 1) x = ‖x‖ := -begin - obtain rfl | hx := eq_or_ne x 0, - { rw [norm_zero, gauge_zero] }, - refine (le_of_forall_pos_le_add $ λ ε hε, _).antisymm _, - { have : 0 < ‖x‖ + ε := by positivity, - refine gauge_le_of_mem this.le _, - rw [smul_ball this.ne', smul_zero, real.norm_of_nonneg this.le, mul_one, mem_ball_zero_iff], - exact lt_add_of_pos_right _ hε }, - refine le_gauge_of_not_mem balanced_ball_zero.star_convex - (absorbent_ball_zero zero_lt_one).absorbs (λ h, _), - obtain hx' | hx' := eq_or_ne (‖x‖) 0, - { rw hx' at h, - exact hx (zero_smul_set_subset _ h) }, - { rw [mem_smul_set_iff_inv_smul_mem₀ hx', mem_ball_zero_iff, norm_smul, norm_inv, norm_norm, - inv_mul_cancel hx'] at h, - exact lt_irrefl _ h } -end +by rw [← ball_norm_seminorm ℝ, seminorm.gauge_ball, coe_norm_seminorm] lemma gauge_ball (hr : 0 < r) (x : E) : gauge (metric.ball (0 : E) r) x = ‖x‖ / r := begin @@ -472,4 +456,23 @@ begin exact gauge_mono (absorbent_ball_zero hr) hs x, end +lemma convex.lipschitz_with_gauge {r : ℝ≥0} (hc : convex ℝ s) (hr : 0 < r) + (hs : metric.ball (0 : E) r ⊆ s) : + lipschitz_with r⁻¹ (gauge s) := +have absorbent ℝ (metric.ball (0 : E) r) := absorbent_ball_zero hr, +lipschitz_with.of_le_add_mul _ $ λ x y, + calc gauge s x = gauge s (y + (x - y)) : by simp + ... ≤ gauge s y + gauge s (x - y) : gauge_add_le hc (this.subset hs) _ _ + ... ≤ gauge s y + ‖x - y‖ / r : + add_le_add_left ((gauge_mono this hs (x - y)).trans_eq (gauge_ball hr _)) _ + ... = gauge s y + r⁻¹ * dist x y : by rw [dist_eq_norm, div_eq_inv_mul] + +lemma convex.uniform_continuous_gauge (hc : convex ℝ s) (h₀ : s ∈ 𝓝 (0 : E)) : + uniform_continuous (gauge s) := +begin + obtain ⟨r, hr₀, hr⟩ := metric.mem_nhds_iff.1 h₀, + lift r to ℝ≥0 using le_of_lt hr₀, + exact (hc.lipschitz_with_gauge hr₀ hr).uniform_continuous +end + end norm From 62748956a1ece9b26b33243e2e3a2852176666f5 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 26 May 2023 16:13:46 +0000 Subject: [PATCH 1091/1291] fix: missing continuity attribute (#19104) --- src/analysis/normed_space/exponential.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/analysis/normed_space/exponential.lean b/src/analysis/normed_space/exponential.lean index c75855885cbb5..b0d62cf8427b4 100644 --- a/src/analysis/normed_space/exponential.lean +++ b/src/analysis/normed_space/exponential.lean @@ -412,7 +412,7 @@ lemma exp_has_fpower_series_at_zero : has_fpower_series_at (exp 𝕂) (exp_series 𝕂 𝔸) 0 := exp_has_fpower_series_on_ball.has_fpower_series_at -lemma exp_continuous : continuous (exp 𝕂 : 𝔸 → 𝔸) := +@[continuity] lemma exp_continuous : continuous (exp 𝕂 : 𝔸 → 𝔸) := begin rw [continuous_iff_continuous_on_univ, ← metric.eball_top_eq_univ (0 : 𝔸), ← exp_series_radius_eq_top 𝕂 𝔸], From ea0bcd84221246c801a6f8fbe8a4372f6d04b176 Mon Sep 17 00:00:00 2001 From: "Filippo A. E. Nuccio" Date: Fri, 26 May 2023 19:59:45 +0000 Subject: [PATCH 1092/1291] feat(number_theory/class_number/finite): remove useless assumption [dedekind_domain R] (#19109) Remove the useless assumption `[dedekind_domain R]` from the final lemma [class_group.fintype_of_admissible_of_finite](https://leanprover-community.github.io/mathlib_docs/number_theory/class_number/finite.html#class_group.fintype_of_admissible_of_finite) since this can be inferred by the Euclidean domain one, that is in force thanks the standing assumption is that `R` has an admissible absolute value. --- src/number_theory/class_number/finite.lean | 3 +-- 1 file changed, 1 insertion(+), 2 deletions(-) diff --git a/src/number_theory/class_number/finite.lean b/src/number_theory/class_number/finite.lean index fb56d22dbdbb3..3e868fbb2d3b7 100644 --- a/src/number_theory/class_number/finite.lean +++ b/src/number_theory/class_number/finite.lean @@ -378,8 +378,7 @@ absolute value. See also `class_group.fintype_of_admissible_of_algebraic` where `L` is an algebraic extension of `R`, that includes some extra assumptions. -/ -noncomputable def fintype_of_admissible_of_finite [is_dedekind_domain R] : - fintype (class_group S) := +noncomputable def fintype_of_admissible_of_finite : fintype (class_group S) := begin letI := classical.dec_eq L, letI := is_integral_closure.is_fraction_ring_of_finite_extension R K L S, From b84aee748341da06a6d78491367e2c0e9f15e8a5 Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Sat, 27 May 2023 07:13:41 +0000 Subject: [PATCH 1093/1291] feat (measure_theory/integral): some simple substitution rules (#19106) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This adds a few lemmas for evaluating integrals, and showing that integrability is preserved, by substitutions of the form `x ↦ a * x` or `x ↦ x ^ a`. --- src/measure_theory/function/jacobian.lean | 39 +++++++--- .../integral/integral_eq_improper.lean | 72 ++++++++++++++++++ .../measure/haar/normed_space.lean | 75 +++++++++++++++++++ 3 files changed, 174 insertions(+), 12 deletions(-) diff --git a/src/measure_theory/function/jacobian.lean b/src/measure_theory/function/jacobian.lean index 52a765f701dd1..ce5efac87502b 100644 --- a/src/measure_theory/function/jacobian.lean +++ b/src/measure_theory/function/jacobian.lean @@ -1245,24 +1245,39 @@ begin refl end +/- Porting note: move this to `topology.algebra.module.basic` when port is over -/ +lemma det_one_smul_right {𝕜 : Type*} [normed_field 𝕜] (v : 𝕜) : + ((1 : 𝕜 →L[𝕜] 𝕜).smul_right v).det = v := +begin + have : (1 : 𝕜 →L[𝕜] 𝕜).smul_right v = v • (1 : 𝕜 →L[𝕜] 𝕜), + { ext1 w, + simp only [continuous_linear_map.smul_right_apply, continuous_linear_map.one_apply, + algebra.id.smul_eq_mul, one_mul, continuous_linear_map.coe_smul', pi.smul_apply, mul_one] }, + rw [this, continuous_linear_map.det, continuous_linear_map.coe_smul], + change ((1 : 𝕜 →L[𝕜] 𝕜) : 𝕜 →ₗ[𝕜] 𝕜) with linear_map.id, + rw [linear_map.det_smul, finite_dimensional.finrank_self, linear_map.det_id, pow_one, mul_one], +end + +/-- Integrability in the change of variable formula for differentiable functions (one-variable +version): if a function `f` is injective and differentiable on a measurable set ``s ⊆ ℝ`, then a +function `g : ℝ → F` is integrable on `f '' s` if and only if `|(f' x)| • g ∘ f` is integrable on +`s`. -/ +theorem integrable_on_image_iff_integrable_on_abs_deriv_smul + {s : set ℝ} {f : ℝ → ℝ} {f' : ℝ → ℝ} (hs : measurable_set s) + (hf' : ∀ x ∈ s, has_deriv_within_at f (f' x) s x) (hf : inj_on f s) (g : ℝ → F) : + integrable_on g (f '' s) ↔ integrable_on (λ x, |(f' x)| • g (f x)) s := +by simpa only [det_one_smul_right] using integrable_on_image_iff_integrable_on_abs_det_fderiv_smul + volume hs (λ x hx, (hf' x hx).has_fderiv_within_at) hf g + /-- Change of variable formula for differentiable functions (one-variable version): if a function `f` is injective and differentiable on a measurable set `s ⊆ ℝ`, then the Bochner integral of a -function `g : ℝ → F` on `f '' s` coincides with the integral of `|(f' x).det| • g ∘ f` on `s`. -/ +function `g : ℝ → F` on `f '' s` coincides with the integral of `|(f' x)| • g ∘ f` on `s`. -/ theorem integral_image_eq_integral_abs_deriv_smul {s : set ℝ} {f : ℝ → ℝ} {f' : ℝ → ℝ} [complete_space F] (hs : measurable_set s) (hf' : ∀ x ∈ s, has_deriv_within_at f (f' x) s x) (hf : inj_on f s) (g : ℝ → F) : ∫ x in f '' s, g x = ∫ x in s, |(f' x)| • g (f x) := -begin - convert integral_image_eq_integral_abs_det_fderiv_smul volume hs - (λ x hx, (hf' x hx).has_fderiv_within_at) hf g, - ext1 x, - rw (by { ext, simp } : (1 : ℝ →L[ℝ] ℝ).smul_right (f' x) = (f' x) • (1 : ℝ →L[ℝ] ℝ)), - rw [continuous_linear_map.det, continuous_linear_map.coe_smul], - have : ((1 : ℝ →L[ℝ] ℝ) : ℝ →ₗ[ℝ] ℝ) = (1 : ℝ →ₗ[ℝ] ℝ) := by refl, - rw [this, linear_map.det_smul, finite_dimensional.finrank_self], - suffices : (1 : ℝ →ₗ[ℝ] ℝ).det = 1, { rw this, simp }, - exact linear_map.det_id, -end +by simpa only [det_one_smul_right] using integral_image_eq_integral_abs_det_fderiv_smul + volume hs (λ x hx, (hf' x hx).has_fderiv_within_at) hf g theorem integral_target_eq_integral_abs_det_fderiv_smul [complete_space F] {f : local_homeomorph E E} (hf' : ∀ x ∈ f.source, has_fderiv_at f (f' x) x) (g : E → F) : diff --git a/src/measure_theory/integral/integral_eq_improper.lean b/src/measure_theory/integral/integral_eq_improper.lean index b46ea33b9b080..7268b5289486b 100644 --- a/src/measure_theory/integral/integral_eq_improper.lean +++ b/src/measure_theory/integral/integral_eq_improper.lean @@ -7,6 +7,7 @@ import analysis.special_functions.pow.deriv import measure_theory.integral.fund_thm_calculus import order.filter.at_top_bot import measure_theory.function.jacobian +import measure_theory.measure.haar.normed_space /-! # Links between an integral and its "improper" version @@ -914,6 +915,77 @@ begin funext, congr, rw abs_of_nonneg hp.le, end +lemma integral_comp_mul_left_Ioi (g : ℝ → E) (a : ℝ) {b : ℝ} (hb : 0 < b) : + ∫ x in Ioi a, g (b * x) = |b⁻¹| • ∫ x in Ioi (b * a), g x := +begin + have : ∀ c : ℝ, measurable_set (Ioi c) := λ c, measurable_set_Ioi, + rw [←integral_indicator (this a), ←integral_indicator (this $ b * a)], + convert measure.integral_comp_mul_left _ b, + ext1 x, + rw [←indicator_comp_right, preimage_const_mul_Ioi _ hb, mul_div_cancel_left _ hb.ne'], +end + +lemma integral_comp_mul_right_Ioi (g : ℝ → E) (a : ℝ) {b : ℝ} (hb : 0 < b) : + ∫ x in Ioi a, g (x * b) = |b⁻¹| • ∫ x in Ioi (a * b), g x := +by simpa only [mul_comm] using integral_comp_mul_left_Ioi g a hb + end Ioi_change_variables +section Ioi_integrability + +open real +open_locale interval + +variables {E : Type*} [normed_add_comm_group E] + +/-- The substitution `y = x ^ p` in integrals over `Ioi 0` preserves integrability. -/ +lemma integrable_on_Ioi_comp_rpow_iff [normed_space ℝ E] (f : ℝ → E) {p : ℝ} (hp : p ≠ 0) : + integrable_on (λ x, (|p| * x ^ (p - 1)) • f (x ^ p)) (Ioi 0) ↔ integrable_on f (Ioi 0) := +begin + let S := Ioi (0 : ℝ), + have a1 : ∀ x:ℝ, x ∈ S → has_deriv_within_at (λ (t:ℝ), t ^ p) (p * x ^ (p - 1)) S x := + λ x hx, (has_deriv_at_rpow_const (or.inl (mem_Ioi.mp hx).ne')).has_deriv_within_at, + have a2 : inj_on (λ x:ℝ, x ^ p) S, + { rcases lt_or_gt_of_ne hp, + { apply strict_anti_on.inj_on, + intros x hx y hy hxy, + rw [←inv_lt_inv (rpow_pos_of_pos hx p) (rpow_pos_of_pos hy p), + ←rpow_neg (le_of_lt hx), ←rpow_neg (le_of_lt hy)], + exact rpow_lt_rpow (le_of_lt hx) hxy (neg_pos.mpr h), }, + exact strict_mono_on.inj_on (λ x hx y hy hxy, rpow_lt_rpow (mem_Ioi.mp hx).le hxy h) }, + have a3 : (λ (t : ℝ), t ^ p) '' S = S, + { ext1, rw mem_image, split, + { rintro ⟨y, hy, rfl⟩, exact rpow_pos_of_pos hy p }, + { intro hx, refine ⟨x ^ (1 / p), rpow_pos_of_pos hx _, _⟩, + rw [←rpow_mul (le_of_lt hx), one_div_mul_cancel hp, rpow_one], } }, + have := integrable_on_image_iff_integrable_on_abs_deriv_smul measurable_set_Ioi a1 a2 f, + rw a3 at this, + rw this, + refine integrable_on_congr_fun (λ x hx, _) measurable_set_Ioi, + simp_rw [abs_mul, abs_of_nonneg (rpow_nonneg_of_nonneg (le_of_lt hx) _)], +end + +/-- The substitution `y = x ^ p` in integrals over `Ioi 0` preserves integrability (version +without `|p|` factor) -/ +lemma integrable_on_Ioi_comp_rpow_iff' [normed_space ℝ E] (f : ℝ → E) {p : ℝ} (hp : p ≠ 0) : + integrable_on (λ x, x ^ (p - 1) • f (x ^ p)) (Ioi 0) ↔ integrable_on f (Ioi 0) := +by simpa only [←integrable_on_Ioi_comp_rpow_iff f hp, mul_smul] + using (integrable_smul_iff (abs_pos.mpr hp).ne' _).symm + +lemma integrable_on_Ioi_comp_mul_left_iff (f : ℝ → E) (c : ℝ) {a : ℝ} (ha : 0 < a) : + integrable_on (λ x, f (a * x)) (Ioi c) ↔ integrable_on f (Ioi $ a * c) := +begin + rw [←integrable_indicator_iff (measurable_set_Ioi : measurable_set $ Ioi c)], + rw [←integrable_indicator_iff (measurable_set_Ioi : measurable_set $ Ioi $ a * c)], + convert integrable_comp_mul_left_iff ((Ioi (a * c)).indicator f) ha.ne' using 2, + ext1 x, + rw [←indicator_comp_right, preimage_const_mul_Ioi _ ha, mul_comm a c, mul_div_cancel _ ha.ne'], +end + +lemma integrable_on_Ioi_comp_mul_right_iff (f : ℝ → E) (c : ℝ) {a : ℝ} (ha : 0 < a) : + integrable_on (λ x, f (x * a)) (Ioi c) ↔ integrable_on f (Ioi $ c * a) := +by simpa only [mul_comm, mul_zero] using integrable_on_Ioi_comp_mul_left_iff f c ha + +end Ioi_integrability + end measure_theory diff --git a/src/measure_theory/measure/haar/normed_space.lean b/src/measure_theory/measure/haar/normed_space.lean index 5fe6c03c96386..49aae295bb8dd 100644 --- a/src/measure_theory/measure/haar/normed_space.lean +++ b/src/measure_theory/measure/haar/normed_space.lean @@ -98,5 +98,80 @@ lemma integral_comp_inv_smul_of_nonneg (f : E → F) {R : ℝ} (hR : 0 ≤ R) : ∫ x, f (R⁻¹ • x) ∂μ = R ^ finrank ℝ E • ∫ x, f x ∂μ := by rw [integral_comp_inv_smul μ f R, abs_of_nonneg ((pow_nonneg hR _))] +lemma integral_comp_mul_left (g : ℝ → F) (a : ℝ) : + ∫ x : ℝ, g (a * x) = |a⁻¹| • ∫ y : ℝ, g y := +by simp_rw [←smul_eq_mul, measure.integral_comp_smul, finite_dimensional.finrank_self, pow_one] + +lemma integral_comp_inv_mul_left (g : ℝ → F) (a : ℝ) : + ∫ x : ℝ, g (a⁻¹ * x) = |a| • ∫ y : ℝ, g y := +by simp_rw [←smul_eq_mul, measure.integral_comp_inv_smul, finite_dimensional.finrank_self, pow_one] + +lemma integral_comp_mul_right (g : ℝ → F) (a : ℝ) : + ∫ x : ℝ, g (x * a) = |a⁻¹| • ∫ y : ℝ, g y := +by simpa only [mul_comm] using integral_comp_mul_left g a + +lemma integral_comp_inv_mul_right (g : ℝ → F) (a : ℝ) : + ∫ x : ℝ, g (x * a⁻¹) = |a| • ∫ y : ℝ, g y := +by simpa only [mul_comm] using integral_comp_inv_mul_left g a + +lemma integral_comp_div (g : ℝ → F) (a : ℝ) : + ∫ x : ℝ, g (x / a) = |a| • ∫ y : ℝ, g y := +integral_comp_inv_mul_right g a + end measure + +variables {F : Type*} [normed_add_comm_group F] + +lemma integrable_comp_smul_iff {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] + [measurable_space E] [borel_space E] [finite_dimensional ℝ E] + (μ : measure E) [is_add_haar_measure μ] (f : E → F) {R : ℝ} (hR : R ≠ 0) : + integrable (λ x, f (R • x)) μ ↔ integrable f μ := +begin + -- reduce to one-way implication + suffices : ∀ {g : E → F} (hg : integrable g μ) {S : ℝ} (hS : S ≠ 0), + integrable (λ x, g (S • x)) μ, + { refine ⟨λ hf, _, λ hf, this hf hR⟩, + convert this hf (inv_ne_zero hR), + ext1 x, + rw [←mul_smul, mul_inv_cancel hR, one_smul], }, + -- now prove + intros g hg S hS, + let t := ((homeomorph.smul (is_unit_iff_ne_zero.2 hS).unit).to_measurable_equiv : E ≃ᵐ E), + refine (integrable_map_equiv t g).mp (_ : integrable g (map (has_smul.smul S) μ)), + rwa [map_add_haar_smul μ hS, integrable_smul_measure _ ennreal.of_real_ne_top], + simpa only [ne.def, ennreal.of_real_eq_zero, not_le, abs_pos] + using inv_ne_zero (pow_ne_zero _ hS), +end + +lemma integrable.comp_smul {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] + [measurable_space E] [borel_space E] [finite_dimensional ℝ E] + {μ : measure E} [is_add_haar_measure μ] {f : E → F} (hf : integrable f μ) {R : ℝ} (hR : R ≠ 0) : + integrable (λ x, f (R • x)) μ := +(integrable_comp_smul_iff μ f hR).2 hf + +lemma integrable_comp_mul_left_iff (g : ℝ → F) {R : ℝ} (hR : R ≠ 0) : + integrable (λ x, g (R * x)) ↔ integrable g := +by simpa only [smul_eq_mul] using integrable_comp_smul_iff volume g hR + +lemma integrable.comp_mul_left' {g : ℝ → F} (hg : integrable g) {R : ℝ} (hR : R ≠ 0) : + integrable (λ x, g (R * x)) := +(integrable_comp_mul_left_iff g hR).2 hg + +lemma integrable_comp_mul_right_iff (g : ℝ → F) {R : ℝ} (hR : R ≠ 0) : + integrable (λ x, g (x * R)) ↔ integrable g := +by simpa only [mul_comm] using integrable_comp_mul_left_iff g hR + +lemma integrable.comp_mul_right' {g : ℝ → F} (hg : integrable g) {R : ℝ} (hR : R ≠ 0) : + integrable (λ x, g (x * R)) := +(integrable_comp_mul_right_iff g hR).2 hg + +lemma integrable_comp_div_iff (g : ℝ → F) {R : ℝ} (hR : R ≠ 0) : + integrable (λ x, g (x / R)) ↔ integrable g := +integrable_comp_mul_right_iff g (inv_ne_zero hR) + +lemma integrable.comp_div {g : ℝ → F} (hg : integrable g) {R : ℝ} (hR : R ≠ 0) : + integrable (λ x, g (x / R)) := +(integrable_comp_div_iff g hR).2 hg + + end measure_theory From 78fdf68dcd2fdb3fe64c0dd6f88926a49418a6ea Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sat, 27 May 2023 08:18:54 +0000 Subject: [PATCH 1094/1291] feat(linear_algebra/alternating): add 3 missing definitions (#19069) --- src/linear_algebra/alternating.lean | 50 +++++++++++++++++++++++ src/linear_algebra/multilinear/basic.lean | 2 +- 2 files changed, 51 insertions(+), 1 deletion(-) diff --git a/src/linear_algebra/alternating.lean b/src/linear_algebra/alternating.lean index 56101ce060757..c0a1adb614edd 100644 --- a/src/linear_algebra/alternating.lean +++ b/src/linear_algebra/alternating.lean @@ -51,6 +51,7 @@ using `map_swap` as a definition, and does not require `has_neg N`. variables {R : Type*} [semiring R] variables {M : Type*} [add_comm_monoid M] [module R M] variables {N : Type*} [add_comm_monoid N] [module R N] +variables {P : Type*} [add_comm_monoid P] [module R P] -- semiring / add_comm_group variables {M' : Type*} [add_comm_group M'] [module R M'] @@ -207,6 +208,49 @@ instance [distrib_mul_action Sᵐᵒᵖ N] [is_central_scalar S N] : end has_smul +/-- The cartesian product of two alternating maps, as a multilinear map. -/ +@[simps { simp_rhs := tt }] +def prod (f : alternating_map R M N ι) (g : alternating_map R M P ι) : + alternating_map R M (N × P) ι := +{ map_eq_zero_of_eq' := λ v i j h hne, prod.ext (f.map_eq_zero_of_eq _ h hne) + (g.map_eq_zero_of_eq _ h hne), + .. f.to_multilinear_map.prod g.to_multilinear_map } + +@[simp] +lemma coe_prod (f : alternating_map R M N ι) (g : alternating_map R M P ι) : + (f.prod g : multilinear_map R (λ _ : ι, M) (N × P)) = multilinear_map.prod f g := +rfl + +/-- Combine a family of alternating maps with the same domain and codomains `N i` into an +alternating map taking values in the space of functions `Π i, N i`. -/ +@[simps { simp_rhs := tt }] +def pi {ι' : Type*} {N : ι' → Type*} [∀ i, add_comm_monoid (N i)] [∀ i, module R (N i)] + (f : ∀ i, alternating_map R M (N i) ι) : alternating_map R M (∀ i, N i) ι := +{ map_eq_zero_of_eq' := λ v i j h hne, funext $ λ a, (f a).map_eq_zero_of_eq _ h hne, + .. multilinear_map.pi (λ a, (f a).to_multilinear_map) } + +@[simp] +lemma coe_pi {ι' : Type*} {N : ι' → Type*} [∀ i, add_comm_monoid (N i)] + [∀ i, module R (N i)] (f : ∀ i, alternating_map R M (N i) ι) : + (pi f : multilinear_map R (λ _ : ι, M) (∀ i, N i)) = multilinear_map.pi (λ a, f a) := +rfl + +/-- Given an alternating `R`-multilinear map `f` taking values in `R`, `f.smul_right z` is the map +sending `m` to `f m • z`. -/ +@[simps { simp_rhs := tt }] +def smul_right {R M₁ M₂ ι : Type*} [comm_semiring R] + [add_comm_monoid M₁] [add_comm_monoid M₂] [module R M₁] [module R M₂] + (f : alternating_map R M₁ R ι) (z : M₂) : alternating_map R M₁ M₂ ι := +{ map_eq_zero_of_eq' := λ v i j h hne, by simp [f.map_eq_zero_of_eq v h hne], + .. f.to_multilinear_map.smul_right z } + +@[simp] +lemma coe_smul_right {R M₁ M₂ ι : Type*} [comm_semiring R] + [add_comm_monoid M₁] [add_comm_monoid M₂] [module R M₁] [module R M₂] + (f : alternating_map R M₁ R ι) (z : M₂) : + (f.smul_right z : multilinear_map R (λ _ : ι, M₁) M₂) = multilinear_map.smul_right f z := +rfl + instance : has_add (alternating_map R M N ι) := ⟨λ a b, { map_eq_zero_of_eq' := @@ -335,6 +379,12 @@ def comp_alternating_map (g : N →ₗ[R] N₂) : alternating_map R M N ι →+ lemma comp_alternating_map_apply (g : N →ₗ[R] N₂) (f : alternating_map R M N ι) (m : ι → M) : g.comp_alternating_map f m = g (f m) := rfl +lemma smul_right_eq_comp {R M₁ M₂ ι : Type*} [comm_semiring R] + [add_comm_monoid M₁] [add_comm_monoid M₂] [module R M₁] [module R M₂] + (f : alternating_map R M₁ R ι) (z : M₂) : + f.smul_right z = (linear_map.id.smul_right z).comp_alternating_map f := +rfl + @[simp] lemma subtype_comp_alternating_map_cod_restrict (f : alternating_map R M N ι) (p : submodule R N) (h) : diff --git a/src/linear_algebra/multilinear/basic.lean b/src/linear_algebra/multilinear/basic.lean index bef802e1d31f6..1bc50a50377a7 100644 --- a/src/linear_algebra/multilinear/basic.lean +++ b/src/linear_algebra/multilinear/basic.lean @@ -206,7 +206,7 @@ coordinates but `i` equal to those of `m`, and varying the `i`-th coordinate. -/ map_smul' := λc x, by simp } /-- The cartesian product of two multilinear maps, as a multilinear map. -/ -def prod (f : multilinear_map R M₁ M₂) (g : multilinear_map R M₁ M₃) : +@[simps] def prod (f : multilinear_map R M₁ M₂) (g : multilinear_map R M₁ M₃) : multilinear_map R M₁ (M₂ × M₃) := { to_fun := λ m, (f m, g m), map_add' := λ _ m i x y, by simp, From c4015acc0a223449d44061e27ddac1835a3852b9 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sat, 27 May 2023 11:16:23 +0000 Subject: [PATCH 1095/1291] refactor(measure_theory/function/lp_space): split file (#19112) This is the longest file that remains to port, and there is an obvious split to be made. The new file is called `measure_theory.function/lp_seminorm`. Other than the module docstrings, which have been tweaked to represent the split, the contents of the new file is moved without modification from the old one. --- src/measure_theory/function/lp_seminorm.lean | 1630 ++++++++++++++++++ src/measure_theory/function/lp_space.lean | 1612 +---------------- 2 files changed, 1635 insertions(+), 1607 deletions(-) create mode 100644 src/measure_theory/function/lp_seminorm.lean diff --git a/src/measure_theory/function/lp_seminorm.lean b/src/measure_theory/function/lp_seminorm.lean new file mode 100644 index 0000000000000..54e5e97930327 --- /dev/null +++ b/src/measure_theory/function/lp_seminorm.lean @@ -0,0 +1,1630 @@ +/- +Copyright (c) 2020 Rémy Degenne. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Rémy Degenne, Sébastien Gouëzel +-/ +import analysis.normed_space.indicator_function +import analysis.special_functions.pow.continuity +import measure_theory.function.ess_sup +import measure_theory.function.ae_eq_fun +import measure_theory.integral.mean_inequalities + +/-! +# ℒp space + +This file describes properties of almost everywhere strongly measurable functions with finite +`p`-seminorm, denoted by `snorm f p μ` and defined for `p:ℝ≥0∞` as `0` if `p=0`, +`(∫ ‖f a‖^p ∂μ) ^ (1/p)` for `0 < p < ∞` and `ess_sup ‖f‖ μ` for `p=∞`. + +The Prop-valued `mem_ℒp f p μ` states that a function `f : α → E` has finite `p`-seminorm +and is almost everywhere strongly measurable. + +## Main definitions + +* `snorm' f p μ` : `(∫ ‖f a‖^p ∂μ) ^ (1/p)` for `f : α → F` and `p : ℝ`, where `α` is a measurable + space and `F` is a normed group. +* `snorm_ess_sup f μ` : seminorm in `ℒ∞`, equal to the essential supremum `ess_sup ‖f‖ μ`. +* `snorm f p μ` : for `p : ℝ≥0∞`, seminorm in `ℒp`, equal to `0` for `p=0`, to `snorm' f p μ` + for `0 < p < ∞` and to `snorm_ess_sup f μ` for `p = ∞`. + +* `mem_ℒp f p μ` : property that the function `f` is almost everywhere strongly measurable and has + finite `p`-seminorm for the measure `μ` (`snorm f p μ < ∞`) + +-/ + +noncomputable theory +open topological_space measure_theory filter +open_locale nnreal ennreal big_operators topology measure_theory + +variables {α E F G : Type*} {m m0 : measurable_space α} {p : ℝ≥0∞} {q : ℝ} {μ ν : measure α} + [normed_add_comm_group E] [normed_add_comm_group F] [normed_add_comm_group G] + +namespace measure_theory + +section ℒp + +/-! +### ℒp seminorm + +We define the ℒp seminorm, denoted by `snorm f p μ`. For real `p`, it is given by an integral +formula (for which we use the notation `snorm' f p μ`), and for `p = ∞` it is the essential +supremum (for which we use the notation `snorm_ess_sup f μ`). + +We also define a predicate `mem_ℒp f p μ`, requesting that a function is almost everywhere +strongly measurable and has finite `snorm f p μ`. + +This paragraph is devoted to the basic properties of these definitions. It is constructed as +follows: for a given property, we prove it for `snorm'` and `snorm_ess_sup` when it makes sense, +deduce it for `snorm`, and translate it in terms of `mem_ℒp`. +-/ + +section ℒp_space_definition + +/-- `(∫ ‖f a‖^q ∂μ) ^ (1/q)`, which is a seminorm on the space of measurable functions for which +this quantity is finite -/ +def snorm' {m : measurable_space α} (f : α → F) (q : ℝ) (μ : measure α) : ℝ≥0∞ := +(∫⁻ a, ‖f a‖₊^q ∂μ) ^ (1/q) + +/-- seminorm for `ℒ∞`, equal to the essential supremum of `‖f‖`. -/ +def snorm_ess_sup {m : measurable_space α} (f : α → F) (μ : measure α) := +ess_sup (λ x, (‖f x‖₊ : ℝ≥0∞)) μ + +/-- `ℒp` seminorm, equal to `0` for `p=0`, to `(∫ ‖f a‖^p ∂μ) ^ (1/p)` for `0 < p < ∞` and to +`ess_sup ‖f‖ μ` for `p = ∞`. -/ +def snorm {m : measurable_space α} (f : α → F) (p : ℝ≥0∞) (μ : measure α) : ℝ≥0∞ := +if p = 0 then 0 else (if p = ∞ then snorm_ess_sup f μ else snorm' f (ennreal.to_real p) μ) + +lemma snorm_eq_snorm' (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) {f : α → F} : + snorm f p μ = snorm' f (ennreal.to_real p) μ := +by simp [snorm, hp_ne_zero, hp_ne_top] + +lemma snorm_eq_lintegral_rpow_nnnorm (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) {f : α → F} : + snorm f p μ = (∫⁻ x, ‖f x‖₊ ^ p.to_real ∂μ) ^ (1 / p.to_real) := +by rw [snorm_eq_snorm' hp_ne_zero hp_ne_top, snorm'] + +lemma snorm_one_eq_lintegral_nnnorm {f : α → F} : snorm f 1 μ = ∫⁻ x, ‖f x‖₊ ∂μ := +by simp_rw [snorm_eq_lintegral_rpow_nnnorm one_ne_zero ennreal.coe_ne_top, ennreal.one_to_real, + one_div_one, ennreal.rpow_one] + +@[simp] lemma snorm_exponent_top {f : α → F} : snorm f ∞ μ = snorm_ess_sup f μ := by simp [snorm] + +/-- 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 α . 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 μ) : + ae_strongly_measurable f μ := h.1 + +lemma lintegral_rpow_nnnorm_eq_rpow_snorm' {f : α → F} (hq0_lt : 0 < q) : + ∫⁻ a, ‖f a‖₊ ^ q ∂μ = (snorm' f q μ) ^ q := +begin + rw [snorm', ←ennreal.rpow_mul, one_div, inv_mul_cancel, ennreal.rpow_one], + exact (ne_of_lt hq0_lt).symm, +end + +end ℒp_space_definition + +section top + +lemma mem_ℒp.snorm_lt_top {f : α → E} (hfp : mem_ℒp f p μ) : snorm f p μ < ∞ := hfp.2 + +lemma mem_ℒp.snorm_ne_top {f : α → E} (hfp : mem_ℒp f p μ) : snorm f p μ ≠ ∞ := ne_of_lt (hfp.2) + +lemma lintegral_rpow_nnnorm_lt_top_of_snorm'_lt_top {f : α → F} (hq0_lt : 0 < q) + (hfq : snorm' f q μ < ∞) : + ∫⁻ a, ‖f a‖₊ ^ q ∂μ < ∞ := +begin + rw lintegral_rpow_nnnorm_eq_rpow_snorm' hq0_lt, + exact ennreal.rpow_lt_top_of_nonneg (le_of_lt hq0_lt) (ne_of_lt hfq), +end + +lemma lintegral_rpow_nnnorm_lt_top_of_snorm_lt_top {f : α → F} (hp_ne_zero : p ≠ 0) + (hp_ne_top : p ≠ ∞) (hfp : snorm f p μ < ∞) : + ∫⁻ a, ‖f a‖₊ ^ p.to_real ∂μ < ∞ := +begin + apply lintegral_rpow_nnnorm_lt_top_of_snorm'_lt_top, + { exact ennreal.to_real_pos hp_ne_zero hp_ne_top }, + { simpa [snorm_eq_snorm' hp_ne_zero hp_ne_top] using hfp } +end + +lemma snorm_lt_top_iff_lintegral_rpow_nnnorm_lt_top {f : α → F} (hp_ne_zero : p ≠ 0) + (hp_ne_top : p ≠ ∞) : + snorm f p μ < ∞ ↔ ∫⁻ a, ‖f a‖₊ ^ p.to_real ∂μ < ∞ := +⟨lintegral_rpow_nnnorm_lt_top_of_snorm_lt_top hp_ne_zero hp_ne_top, + begin + intros h, + have hp' := ennreal.to_real_pos hp_ne_zero hp_ne_top, + have : 0 < 1 / p.to_real := div_pos zero_lt_one hp', + simpa [snorm_eq_lintegral_rpow_nnnorm hp_ne_zero hp_ne_top] using + ennreal.rpow_lt_top_of_nonneg (le_of_lt this) (ne_of_lt h) + end⟩ + +end top + +section zero + +@[simp] lemma snorm'_exponent_zero {f : α → F} : snorm' f 0 μ = 1 := +by rw [snorm', div_zero, ennreal.rpow_zero] + +@[simp] lemma snorm_exponent_zero {f : α → F} : snorm f 0 μ = 0 := +by simp [snorm] + +lemma mem_ℒp_zero_iff_ae_strongly_measurable {f : α → E} : + mem_ℒp f 0 μ ↔ ae_strongly_measurable f μ := +by simp [mem_ℒp, snorm_exponent_zero] + +@[simp] lemma snorm'_zero (hp0_lt : 0 < q) : snorm' (0 : α → F) q μ = 0 := +by simp [snorm', hp0_lt] + +@[simp] lemma snorm'_zero' (hq0_ne : q ≠ 0) (hμ : μ ≠ 0) : snorm' (0 : α → F) q μ = 0 := +begin + cases le_or_lt 0 q with hq0 hq_neg, + { exact snorm'_zero (lt_of_le_of_ne hq0 hq0_ne.symm), }, + { simp [snorm', ennreal.rpow_eq_zero_iff, hμ, hq_neg], }, +end + +@[simp] lemma snorm_ess_sup_zero : snorm_ess_sup (0 : α → F) μ = 0 := +begin + simp_rw [snorm_ess_sup, pi.zero_apply, nnnorm_zero, ennreal.coe_zero, ←ennreal.bot_eq_zero], + exact ess_sup_const_bot, +end + +@[simp] lemma snorm_zero : snorm (0 : α → F) p μ = 0 := +begin + by_cases h0 : p = 0, + { simp [h0], }, + by_cases h_top : p = ∞, + { simp only [h_top, snorm_exponent_top, snorm_ess_sup_zero], }, + rw ←ne.def at h0, + simp [snorm_eq_snorm' h0 h_top, ennreal.to_real_pos h0 h_top], +end + +@[simp] lemma snorm_zero' : snorm (λ x : α, (0 : F)) p μ = 0 := +by convert snorm_zero + +lemma zero_mem_ℒp : mem_ℒp (0 : α → E) p μ := +⟨ae_strongly_measurable_zero, by { rw snorm_zero, exact ennreal.coe_lt_top, } ⟩ + +lemma zero_mem_ℒp' : mem_ℒp (λ x : α, (0 : E)) p μ := +by convert zero_mem_ℒp + +variables [measurable_space α] + +lemma snorm'_measure_zero_of_pos {f : α → F} (hq_pos : 0 < q) : + snorm' f q (0 : measure α) = 0 := +by simp [snorm', hq_pos] + +lemma snorm'_measure_zero_of_exponent_zero {f : α → F} : snorm' f 0 (0 : measure α) = 1 := +by simp [snorm'] + +lemma snorm'_measure_zero_of_neg {f : α → F} (hq_neg : q < 0) : snorm' f q (0 : measure α) = ∞ := +by simp [snorm', hq_neg] + +@[simp] lemma snorm_ess_sup_measure_zero {f : α → F} : snorm_ess_sup f (0 : measure α) = 0 := +by simp [snorm_ess_sup] + +@[simp] lemma snorm_measure_zero {f : α → F} : snorm f p (0 : measure α) = 0 := +begin + by_cases h0 : p = 0, + { simp [h0], }, + by_cases h_top : p = ∞, + { simp [h_top], }, + rw ←ne.def at h0, + simp [snorm_eq_snorm' h0 h_top, snorm', ennreal.to_real_pos h0 h_top], +end + +end zero + +section const + +lemma snorm'_const (c : F) (hq_pos : 0 < q) : + snorm' (λ x : α , c) q μ = (‖c‖₊ : ℝ≥0∞) * (μ set.univ) ^ (1/q) := +begin + rw [snorm', lintegral_const, ennreal.mul_rpow_of_nonneg _ _ (by simp [hq_pos.le] : 0 ≤ 1 / q)], + congr, + rw ←ennreal.rpow_mul, + suffices hq_cancel : q * (1/q) = 1, by rw [hq_cancel, ennreal.rpow_one], + rw [one_div, mul_inv_cancel (ne_of_lt hq_pos).symm], +end + +lemma snorm'_const' [is_finite_measure μ] (c : F) (hc_ne_zero : c ≠ 0) (hq_ne_zero : q ≠ 0) : + snorm' (λ x : α , c) q μ = (‖c‖₊ : ℝ≥0∞) * (μ set.univ) ^ (1/q) := +begin + rw [snorm', lintegral_const, ennreal.mul_rpow_of_ne_top _ (measure_ne_top μ set.univ)], + { congr, + rw ←ennreal.rpow_mul, + suffices hp_cancel : q * (1/q) = 1, by rw [hp_cancel, ennreal.rpow_one], + rw [one_div, mul_inv_cancel hq_ne_zero], }, + { rw [ne.def, ennreal.rpow_eq_top_iff, not_or_distrib, not_and_distrib, not_and_distrib], + split, + { left, + rwa [ennreal.coe_eq_zero, nnnorm_eq_zero], }, + { exact or.inl ennreal.coe_ne_top, }, }, +end + +lemma snorm_ess_sup_const (c : F) (hμ : μ ≠ 0) : + snorm_ess_sup (λ x : α, c) μ = (‖c‖₊ : ℝ≥0∞) := +by rw [snorm_ess_sup, ess_sup_const _ hμ] + +lemma snorm'_const_of_is_probability_measure (c : F) (hq_pos : 0 < q) [is_probability_measure μ] : + snorm' (λ x : α , c) q μ = (‖c‖₊ : ℝ≥0∞) := +by simp [snorm'_const c hq_pos, measure_univ] + +lemma snorm_const (c : F) (h0 : p ≠ 0) (hμ : μ ≠ 0) : + snorm (λ x : α , c) p μ = (‖c‖₊ : ℝ≥0∞) * (μ set.univ) ^ (1/(ennreal.to_real p)) := +begin + by_cases h_top : p = ∞, + { simp [h_top, snorm_ess_sup_const c hμ], }, + simp [snorm_eq_snorm' h0 h_top, snorm'_const, ennreal.to_real_pos h0 h_top], +end + +lemma snorm_const' (c : F) (h0 : p ≠ 0) (h_top: p ≠ ∞) : + snorm (λ x : α , c) p μ = (‖c‖₊ : ℝ≥0∞) * (μ set.univ) ^ (1/(ennreal.to_real p)) := +begin + simp [snorm_eq_snorm' h0 h_top, snorm'_const, ennreal.to_real_pos h0 h_top], +end + +lemma snorm_const_lt_top_iff {p : ℝ≥0∞} {c : F} (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) : + snorm (λ x : α, c) p μ < ∞ ↔ c = 0 ∨ μ set.univ < ∞ := +begin + have hp : 0 < p.to_real, from ennreal.to_real_pos hp_ne_zero hp_ne_top, + by_cases hμ : μ = 0, + { simp only [hμ, measure.coe_zero, pi.zero_apply, or_true, with_top.zero_lt_top, + snorm_measure_zero], }, + by_cases hc : c = 0, + { simp only [hc, true_or, eq_self_iff_true, with_top.zero_lt_top, snorm_zero'], }, + rw snorm_const' c hp_ne_zero hp_ne_top, + by_cases hμ_top : μ set.univ = ∞, + { simp [hc, hμ_top, hp], }, + rw ennreal.mul_lt_top_iff, + simp only [true_and, one_div, ennreal.rpow_eq_zero_iff, hμ, false_or, or_false, + ennreal.coe_lt_top, nnnorm_eq_zero, ennreal.coe_eq_zero, + measure_theory.measure.measure_univ_eq_zero, hp, inv_lt_zero, hc, and_false, false_and, + _root_.inv_pos, or_self, hμ_top, ne.lt_top hμ_top, iff_true], + exact ennreal.rpow_lt_top_of_nonneg (inv_nonneg.mpr hp.le) hμ_top, +end + +lemma mem_ℒp_const (c : E) [is_finite_measure μ] : mem_ℒp (λ a:α, c) p μ := +begin + refine ⟨ae_strongly_measurable_const, _⟩, + by_cases h0 : p = 0, + { simp [h0], }, + by_cases hμ : μ = 0, + { simp [hμ], }, + rw snorm_const c h0 hμ, + refine ennreal.mul_lt_top ennreal.coe_ne_top _, + refine (ennreal.rpow_lt_top_of_nonneg _ (measure_ne_top μ set.univ)).ne, + 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 + rw ← snorm_const_lt_top_iff hp_ne_zero hp_ne_top, + exact ⟨λ h, h.2, λ h, ⟨ae_strongly_measurable_const, h⟩⟩, +end + +end const + +lemma snorm'_mono_nnnorm_ae {f : α → F} {g : α → G} (hq : 0 ≤ q) (h : ∀ᵐ x ∂μ, ‖f x‖₊ ≤ ‖g x‖₊) : + snorm' f q μ ≤ snorm' g q μ := +begin + rw [snorm'], + refine ennreal.rpow_le_rpow _ (one_div_nonneg.2 hq), + refine lintegral_mono_ae (h.mono $ λ x hx, _), + exact ennreal.rpow_le_rpow (ennreal.coe_le_coe.2 hx) hq +end + +lemma snorm'_mono_ae {f : α → F} {g : α → G} (hq : 0 ≤ q) (h : ∀ᵐ x ∂μ, ‖f x‖ ≤ ‖g x‖) : + snorm' f q μ ≤ snorm' g q μ := +snorm'_mono_nnnorm_ae hq h + +lemma snorm'_congr_nnnorm_ae {f g : α → F} (hfg : ∀ᵐ x ∂μ, ‖f x‖₊ = ‖g x‖₊) : + snorm' f q μ = snorm' g q μ := +begin + have : (λ x, (‖f x‖₊ ^ q : ℝ≥0∞)) =ᵐ[μ] (λ x, ‖g x‖₊ ^ q), + from hfg.mono (λ x hx, by simp_rw hx), + simp only [snorm', lintegral_congr_ae this] +end + +lemma snorm'_congr_norm_ae {f g : α → F} (hfg : ∀ᵐ x ∂μ, ‖f x‖ = ‖g x‖) : + snorm' f q μ = snorm' g q μ := +snorm'_congr_nnnorm_ae $ hfg.mono $ λ x hx, nnreal.eq hx + +lemma snorm'_congr_ae {f g : α → F} (hfg : f =ᵐ[μ] g) : snorm' f q μ = snorm' g q μ := +snorm'_congr_nnnorm_ae (hfg.fun_comp _) + +lemma snorm_ess_sup_congr_ae {f g : α → F} (hfg : f =ᵐ[μ] g) : + snorm_ess_sup f μ = snorm_ess_sup g μ := +ess_sup_congr_ae (hfg.fun_comp (coe ∘ nnnorm)) + +lemma snorm_ess_sup_mono_nnnorm_ae {f g : α → F} (hfg : ∀ᵐ x ∂μ, ‖f x‖₊ ≤ ‖g x‖₊) : + snorm_ess_sup f μ ≤ snorm_ess_sup g μ := +ess_sup_mono_ae $ hfg.mono $ λ x hx, ennreal.coe_le_coe.mpr hx + +lemma snorm_mono_nnnorm_ae {f : α → F} {g : α → G} (h : ∀ᵐ x ∂μ, ‖f x‖₊ ≤ ‖g x‖₊) : + snorm f p μ ≤ snorm g p μ := +begin + simp only [snorm], + split_ifs, + { exact le_rfl }, + { exact ess_sup_mono_ae (h.mono $ λ x hx, ennreal.coe_le_coe.mpr hx) }, + { exact snorm'_mono_nnnorm_ae ennreal.to_real_nonneg h } +end + +lemma snorm_mono_ae {f : α → F} {g : α → G} (h : ∀ᵐ x ∂μ, ‖f x‖ ≤ ‖g x‖) : + snorm f p μ ≤ snorm g p μ := +snorm_mono_nnnorm_ae h + +lemma snorm_mono_ae_real {f : α → F} {g : α → ℝ} (h : ∀ᵐ x ∂μ, ‖f x‖ ≤ g x) : + snorm f p μ ≤ snorm g p μ := +snorm_mono_ae $ h.mono (λ x hx, hx.trans ((le_abs_self _).trans (real.norm_eq_abs _).symm.le)) + +lemma snorm_mono_nnnorm {f : α → F} {g : α → G} (h : ∀ x, ‖f x‖₊ ≤ ‖g x‖₊) : + snorm f p μ ≤ snorm g p μ := +snorm_mono_nnnorm_ae (eventually_of_forall (λ x, h x)) + +lemma snorm_mono {f : α → F} {g : α → G} (h : ∀ x, ‖f x‖ ≤ ‖g x‖) : + snorm f p μ ≤ snorm g p μ := +snorm_mono_nnnorm h + +lemma snorm_mono_real {f : α → F} {g : α → ℝ} (h : ∀ x, ‖f x‖ ≤ g x) : + snorm f p μ ≤ snorm g p μ := +snorm_mono_ae_real (eventually_of_forall (λ x, h x)) + +lemma snorm_ess_sup_le_of_ae_nnnorm_bound {f : α → F} {C : ℝ≥0} (hfC : ∀ᵐ x ∂μ, ‖f x‖₊ ≤ C) : + snorm_ess_sup f μ ≤ C := +ess_sup_le_of_ae_le C $ hfC.mono $ λ x hx, ennreal.coe_le_coe.mpr hx + +lemma snorm_ess_sup_le_of_ae_bound {f : α → F} {C : ℝ} (hfC : ∀ᵐ x ∂μ, ‖f x‖ ≤ C) : + snorm_ess_sup f μ ≤ ennreal.of_real C := +snorm_ess_sup_le_of_ae_nnnorm_bound $ hfC.mono $ λ x hx, hx.trans C.le_coe_to_nnreal + +lemma snorm_ess_sup_lt_top_of_ae_nnnorm_bound {f : α → F} {C : ℝ≥0} (hfC : ∀ᵐ x ∂μ, ‖f x‖₊ ≤ C) : + snorm_ess_sup f μ < ∞ := +(snorm_ess_sup_le_of_ae_nnnorm_bound hfC).trans_lt ennreal.coe_lt_top + +lemma snorm_ess_sup_lt_top_of_ae_bound {f : α → F} {C : ℝ} (hfC : ∀ᵐ x ∂μ, ‖f x‖ ≤ C) : + snorm_ess_sup f μ < ∞ := +(snorm_ess_sup_le_of_ae_bound hfC).trans_lt ennreal.of_real_lt_top + +lemma snorm_le_of_ae_nnnorm_bound {f : α → F} {C : ℝ≥0} (hfC : ∀ᵐ x ∂μ, ‖f x‖₊ ≤ C) : + snorm f p μ ≤ C • (μ set.univ ^ p.to_real⁻¹) := +begin + by_cases hμ : μ = 0, + { simp [hμ] }, + haveI : μ.ae.ne_bot := ae_ne_bot.mpr hμ, + by_cases hp : p = 0, + { simp [hp] }, + have : ∀ᵐ x ∂μ, ‖f x‖₊ ≤ ‖(C : ℝ)‖₊ := hfC.mono (λ x hx, hx.trans_eq C.nnnorm_eq.symm), + refine (snorm_mono_ae this).trans_eq _, + rw [snorm_const _ hp hμ, C.nnnorm_eq, one_div, ennreal.smul_def, smul_eq_mul], +end + +lemma snorm_le_of_ae_bound {f : α → F} {C : ℝ} (hfC : ∀ᵐ x ∂μ, ‖f x‖ ≤ C) : + snorm f p μ ≤ ((μ set.univ) ^ p.to_real⁻¹) * (ennreal.of_real C) := +begin + rw [←mul_comm], + exact snorm_le_of_ae_nnnorm_bound (hfC.mono $ λ x hx, hx.trans C.le_coe_to_nnreal), +end + +lemma snorm_congr_nnnorm_ae {f : α → F} {g : α → G} (hfg : ∀ᵐ x ∂μ, ‖f x‖₊ = ‖g x‖₊) : + snorm f p μ = snorm g p μ := +le_antisymm (snorm_mono_nnnorm_ae $ eventually_eq.le hfg) + (snorm_mono_nnnorm_ae $ (eventually_eq.symm hfg).le) + +lemma snorm_congr_norm_ae {f : α → F} {g : α → G} (hfg : ∀ᵐ x ∂μ, ‖f x‖ = ‖g x‖) : + snorm f p μ = snorm g p μ := +snorm_congr_nnnorm_ae $ hfg.mono $ λ x hx, nnreal.eq hx + +@[simp] lemma snorm'_norm {f : α → F} : snorm' (λ a, ‖f a‖) q μ = snorm' f q μ := +by simp [snorm'] + +@[simp] lemma snorm_norm (f : α → F) : snorm (λ x, ‖f x‖) p μ = snorm f p μ := +snorm_congr_norm_ae $ eventually_of_forall $ λ x, norm_norm _ + +lemma snorm'_norm_rpow (f : α → F) (p q : ℝ) (hq_pos : 0 < q) : + snorm' (λ x, ‖f x‖ ^ q) p μ = (snorm' f (p * q) μ) ^ q := +begin + simp_rw snorm', + rw [← ennreal.rpow_mul, ←one_div_mul_one_div], + simp_rw one_div, + rw [mul_assoc, inv_mul_cancel hq_pos.ne.symm, mul_one], + congr, + ext1 x, + simp_rw ← of_real_norm_eq_coe_nnnorm, + rw [real.norm_eq_abs, abs_eq_self.mpr (real.rpow_nonneg_of_nonneg (norm_nonneg _) _), + mul_comm, ← ennreal.of_real_rpow_of_nonneg (norm_nonneg _) hq_pos.le, ennreal.rpow_mul], +end + +lemma snorm_norm_rpow (f : α → F) (hq_pos : 0 < q) : + snorm (λ x, ‖f x‖ ^ q) p μ = (snorm f (p * ennreal.of_real q) μ) ^ q := +begin + by_cases h0 : p = 0, + { simp [h0, ennreal.zero_rpow_of_pos hq_pos], }, + by_cases hp_top : p = ∞, + { simp only [hp_top, snorm_exponent_top, ennreal.top_mul, hq_pos.not_le, ennreal.of_real_eq_zero, + if_false, snorm_exponent_top, snorm_ess_sup], + have h_rpow : ess_sup (λ (x : α), (‖(‖f x‖ ^ q)‖₊ : ℝ≥0∞)) μ + = ess_sup (λ (x : α), (↑‖f x‖₊) ^ q) μ, + { congr, + ext1 x, + nth_rewrite 1 ← nnnorm_norm, + rw [ennreal.coe_rpow_of_nonneg _ hq_pos.le, ennreal.coe_eq_coe], + ext, + push_cast, + rw real.norm_rpow_of_nonneg (norm_nonneg _), }, + rw h_rpow, + have h_rpow_mono := ennreal.strict_mono_rpow_of_pos hq_pos, + have h_rpow_surj := (ennreal.rpow_left_bijective hq_pos.ne.symm).2, + let iso := h_rpow_mono.order_iso_of_surjective _ h_rpow_surj, + exact (iso.ess_sup_apply (λ x, (‖f x‖₊ : ℝ≥0∞)) μ).symm, }, + rw [snorm_eq_snorm' h0 hp_top, snorm_eq_snorm' _ _], + swap, { refine mul_ne_zero h0 _, rwa [ne.def, ennreal.of_real_eq_zero, not_le], }, + swap, { exact ennreal.mul_ne_top hp_top ennreal.of_real_ne_top, }, + rw [ennreal.to_real_mul, ennreal.to_real_of_real hq_pos.le], + exact snorm'_norm_rpow f p.to_real q hq_pos, +end + +lemma snorm_congr_ae {f g : α → F} (hfg : f =ᵐ[μ] g) : snorm f p μ = snorm g p μ := +snorm_congr_norm_ae $ hfg.mono (λ x hx, hx ▸ rfl) + +lemma mem_ℒp_congr_ae {f g : α → E} (hfg : f =ᵐ[μ] g) : mem_ℒp f p μ ↔ mem_ℒp g p μ := +by simp only [mem_ℒp, snorm_congr_ae hfg, ae_strongly_measurable_congr hfg] + +lemma mem_ℒp.ae_eq {f g : α → E} (hfg : f =ᵐ[μ] g) (hf_Lp : mem_ℒp f p μ) : mem_ℒp g p μ := +(mem_ℒp_congr_ae hfg).1 hf_Lp + +lemma mem_ℒp.of_le {f : α → E} {g : α → F} + (hg : mem_ℒp g p μ) (hf : ae_strongly_measurable f μ) (hfg : ∀ᵐ x ∂μ, ‖f x‖ ≤ ‖g x‖) : + mem_ℒp f p μ := +⟨hf, (snorm_mono_ae hfg).trans_lt hg.snorm_lt_top⟩ + +alias mem_ℒp.of_le ← mem_ℒp.mono + +lemma mem_ℒp.mono' {f : α → E} {g : α → ℝ} (hg : mem_ℒp g p μ) + (hf : ae_strongly_measurable f μ) (h : ∀ᵐ a ∂μ, ‖f a‖ ≤ g a) : mem_ℒp f p μ := +hg.mono hf $ h.mono $ λ x hx, le_trans hx (le_abs_self _) + +lemma mem_ℒp.congr_norm {f : α → E} {g : α → F} (hf : mem_ℒp f p μ) + (hg : ae_strongly_measurable g μ) (h : ∀ᵐ a ∂μ, ‖f a‖ = ‖g a‖) : + mem_ℒp g p μ := +hf.mono hg $ eventually_eq.le $ eventually_eq.symm h + +lemma mem_ℒp_congr_norm {f : α → E} {g : α → F} + (hf : ae_strongly_measurable f μ) (hg : ae_strongly_measurable g μ) (h : ∀ᵐ a ∂μ, ‖f a‖ = ‖g a‖) : + mem_ℒp f p μ ↔ mem_ℒp g p μ := +⟨λ h2f, h2f.congr_norm hg h, λ h2g, h2g.congr_norm hf $ eventually_eq.symm h⟩ + +lemma mem_ℒp_top_of_bound {f : α → E} (hf : ae_strongly_measurable f μ) (C : ℝ) + (hfC : ∀ᵐ x ∂μ, ‖f x‖ ≤ C) : + mem_ℒp f ∞ μ := +⟨hf, by { rw snorm_exponent_top, exact snorm_ess_sup_lt_top_of_ae_bound hfC, }⟩ + +lemma mem_ℒp.of_bound [is_finite_measure μ] {f : α → E} (hf : ae_strongly_measurable f μ) + (C : ℝ) (hfC : ∀ᵐ x ∂μ, ‖f x‖ ≤ C) : + mem_ℒp f p μ := +(mem_ℒp_const C).of_le hf (hfC.mono (λ x hx, le_trans hx (le_abs_self _))) + +@[mono] lemma snorm'_mono_measure (f : α → F) (hμν : ν ≤ μ) (hq : 0 ≤ q) : + snorm' f q ν ≤ snorm' f q μ := +begin + simp_rw snorm', + suffices h_integral_mono : (∫⁻ a, (‖f a‖₊ : ℝ≥0∞) ^ q ∂ν) ≤ ∫⁻ a, ‖f a‖₊ ^ q ∂μ, + from ennreal.rpow_le_rpow h_integral_mono (by simp [hq]), + exact lintegral_mono' hμν le_rfl, +end + +@[mono] lemma snorm_ess_sup_mono_measure (f : α → F) (hμν : ν ≪ μ) : + snorm_ess_sup f ν ≤ snorm_ess_sup f μ := +by { simp_rw snorm_ess_sup, exact ess_sup_mono_measure hμν, } + +@[mono] lemma snorm_mono_measure (f : α → F) (hμν : ν ≤ μ) : + snorm f p ν ≤ snorm f p μ := +begin + by_cases hp0 : p = 0, + { simp [hp0], }, + by_cases hp_top : p = ∞, + { simp [hp_top, snorm_ess_sup_mono_measure f (measure.absolutely_continuous_of_le hμν)], }, + simp_rw snorm_eq_snorm' hp0 hp_top, + exact snorm'_mono_measure f hμν ennreal.to_real_nonneg, +end + +lemma mem_ℒp.mono_measure {f : α → E} (hμν : ν ≤ μ) (hf : mem_ℒp f p μ) : + mem_ℒp f p ν := +⟨hf.1.mono_measure hμν, (snorm_mono_measure f hμν).trans_lt hf.2⟩ + +lemma mem_ℒp.restrict (s : set α) {f : α → E} (hf : mem_ℒp f p μ) : + mem_ℒp f p (μ.restrict s) := +hf.mono_measure measure.restrict_le_self + +lemma snorm'_smul_measure {p : ℝ} (hp : 0 ≤ p) {f : α → F} (c : ℝ≥0∞) : + snorm' f p (c • μ) = c ^ (1 / p) * snorm' f p μ := +by { rw [snorm', lintegral_smul_measure, ennreal.mul_rpow_of_nonneg, snorm'], simp [hp], } + +lemma snorm_ess_sup_smul_measure {f : α → F} {c : ℝ≥0∞} (hc : c ≠ 0) : + snorm_ess_sup f (c • μ) = snorm_ess_sup f μ := +by { simp_rw [snorm_ess_sup], exact ess_sup_smul_measure hc, } + +/-- Use `snorm_smul_measure_of_ne_top` instead. -/ +private lemma snorm_smul_measure_of_ne_zero_of_ne_top {p : ℝ≥0∞} (hp_ne_zero : p ≠ 0) + (hp_ne_top : p ≠ ∞) {f : α → F} (c : ℝ≥0∞) : + snorm f p (c • μ) = c ^ (1 / p).to_real • snorm f p μ := +begin + simp_rw snorm_eq_snorm' hp_ne_zero hp_ne_top, + rw snorm'_smul_measure ennreal.to_real_nonneg, + congr, + simp_rw one_div, + rw ennreal.to_real_inv, +end + +lemma snorm_smul_measure_of_ne_zero {p : ℝ≥0∞} {f : α → F} {c : ℝ≥0∞} (hc : c ≠ 0) : + snorm f p (c • μ) = c ^ (1 / p).to_real • snorm f p μ := +begin + by_cases hp0 : p = 0, + { simp [hp0], }, + by_cases hp_top : p = ∞, + { simp [hp_top, snorm_ess_sup_smul_measure hc], }, + exact snorm_smul_measure_of_ne_zero_of_ne_top hp0 hp_top c, +end + +lemma snorm_smul_measure_of_ne_top {p : ℝ≥0∞} (hp_ne_top : p ≠ ∞) {f : α → F} (c : ℝ≥0∞) : + snorm f p (c • μ) = c ^ (1 / p).to_real • snorm f p μ := +begin + by_cases hp0 : p = 0, + { simp [hp0], }, + { exact snorm_smul_measure_of_ne_zero_of_ne_top hp0 hp_ne_top c, }, +end + +lemma snorm_one_smul_measure {f : α → F} (c : ℝ≥0∞) : + snorm f 1 (c • μ) = c * snorm f 1 μ := +by { rw @snorm_smul_measure_of_ne_top _ _ _ μ _ 1 (@ennreal.coe_ne_top 1) f c, simp, } + +lemma mem_ℒp.of_measure_le_smul {μ' : measure α} (c : ℝ≥0∞) (hc : c ≠ ∞) + (hμ'_le : μ' ≤ c • μ) {f : α → E} (hf : mem_ℒp f p μ) : + mem_ℒp f p μ' := +begin + refine ⟨hf.1.mono' (measure.absolutely_continuous_of_le_smul hμ'_le), _⟩, + refine (snorm_mono_measure f hμ'_le).trans_lt _, + by_cases hc0 : c = 0, + { simp [hc0], }, + rw [snorm_smul_measure_of_ne_zero hc0, smul_eq_mul], + refine ennreal.mul_lt_top _ hf.2.ne, + simp [hc, hc0], +end + +lemma mem_ℒp.smul_measure {f : α → E} {c : ℝ≥0∞} (hf : mem_ℒp f p μ) (hc : c ≠ ∞) : + mem_ℒp f p (c • μ) := +hf.of_measure_le_smul c hc le_rfl + +include m + +lemma snorm_one_add_measure (f : α → F) (μ ν : measure α) : + snorm f 1 (μ + ν) = snorm f 1 μ + snorm f 1 ν := +by { simp_rw snorm_one_eq_lintegral_nnnorm, rw lintegral_add_measure _ μ ν, } + +lemma snorm_le_add_measure_right (f : α → F) (μ ν : measure α) {p : ℝ≥0∞} : + snorm f p μ ≤ snorm f p (μ + ν) := +snorm_mono_measure f $ measure.le_add_right $ le_refl _ + +lemma snorm_le_add_measure_left (f : α → F) (μ ν : measure α) {p : ℝ≥0∞} : + snorm f p ν ≤ snorm f p (μ + ν) := +snorm_mono_measure f $ measure.le_add_left $ le_refl _ + +omit m + +lemma mem_ℒp.left_of_add_measure {f : α → E} (h : mem_ℒp f p (μ + ν)) : mem_ℒp f p μ := +h.mono_measure $ measure.le_add_right $ le_refl _ + +lemma mem_ℒp.right_of_add_measure {f : α → E} (h : mem_ℒp f p (μ + ν)) : mem_ℒp f p ν := +h.mono_measure $ measure.le_add_left $ le_refl _ + +lemma mem_ℒp.norm {f : α → E} (h : mem_ℒp f p μ) : mem_ℒp (λ x, ‖f x‖) p μ := +h.of_le h.ae_strongly_measurable.norm (eventually_of_forall (λ x, by simp)) + +lemma mem_ℒp_norm_iff {f : α → E} (hf : ae_strongly_measurable f μ) : + mem_ℒp (λ x, ‖f x‖) p μ ↔ mem_ℒp f p μ := +⟨λ h, ⟨hf, by { rw ← snorm_norm, exact h.2, }⟩, λ h, h.norm⟩ + +lemma snorm'_eq_zero_of_ae_zero {f : α → F} (hq0_lt : 0 < q) (hf_zero : f =ᵐ[μ] 0) : + snorm' f q μ = 0 := +by rw [snorm'_congr_ae hf_zero, snorm'_zero hq0_lt] + +lemma snorm'_eq_zero_of_ae_zero' (hq0_ne : q ≠ 0) (hμ : μ ≠ 0) {f : α → F} (hf_zero : f =ᵐ[μ] 0) : + snorm' f q μ = 0 := +by rw [snorm'_congr_ae hf_zero, snorm'_zero' hq0_ne hμ] + +lemma ae_eq_zero_of_snorm'_eq_zero {f : α → E} (hq0 : 0 ≤ q) (hf : ae_strongly_measurable f μ) + (h : snorm' f q μ = 0) : f =ᵐ[μ] 0 := +begin + rw [snorm', ennreal.rpow_eq_zero_iff] at h, + cases h, + { rw lintegral_eq_zero_iff' (hf.ennnorm.pow_const q) at h, + refine h.left.mono (λ x hx, _), + rw [pi.zero_apply, ennreal.rpow_eq_zero_iff] at hx, + cases hx, + { cases hx with hx _, + rwa [←ennreal.coe_zero, ennreal.coe_eq_coe, nnnorm_eq_zero] at hx, }, + { exact absurd hx.left ennreal.coe_ne_top, }, }, + { exfalso, + rw [one_div, inv_lt_zero] at h, + exact hq0.not_lt h.right }, +end + +lemma snorm'_eq_zero_iff (hq0_lt : 0 < q) {f : α → E} (hf : ae_strongly_measurable f μ) : + snorm' f q μ = 0 ↔ f =ᵐ[μ] 0 := +⟨ae_eq_zero_of_snorm'_eq_zero (le_of_lt hq0_lt) hf, snorm'_eq_zero_of_ae_zero hq0_lt⟩ + +lemma coe_nnnorm_ae_le_snorm_ess_sup {m : measurable_space α} (f : α → F) (μ : measure α) : + ∀ᵐ x ∂μ, (‖f x‖₊ : ℝ≥0∞) ≤ snorm_ess_sup f μ := +ennreal.ae_le_ess_sup (λ x, (‖f x‖₊ : ℝ≥0∞)) + +@[simp] lemma snorm_ess_sup_eq_zero_iff {f : α → F} : snorm_ess_sup f μ = 0 ↔ f =ᵐ[μ] 0 := +by simp [eventually_eq, snorm_ess_sup] + +lemma snorm_eq_zero_iff {f : α → E} (hf : ae_strongly_measurable f μ) (h0 : p ≠ 0) : + snorm f p μ = 0 ↔ f =ᵐ[μ] 0 := +begin + by_cases h_top : p = ∞, + { rw [h_top, snorm_exponent_top, snorm_ess_sup_eq_zero_iff], }, + rw snorm_eq_snorm' h0 h_top, + exact snorm'_eq_zero_iff (ennreal.to_real_pos h0 h_top) hf, +end + +lemma snorm'_add_le {f g : α → E} + (hf : ae_strongly_measurable f μ) (hg : ae_strongly_measurable g μ) (hq1 : 1 ≤ q) : + snorm' (f + g) q μ ≤ snorm' f q μ + snorm' g q μ := +calc (∫⁻ a, ↑‖(f + g) a‖₊ ^ q ∂μ) ^ (1 / q) + ≤ (∫⁻ a, (((λ a, (‖f a‖₊ : ℝ≥0∞)) + + (λ a, (‖g a‖₊ : ℝ≥0∞))) a) ^ q ∂μ) ^ (1 / q) : +begin + refine ennreal.rpow_le_rpow _ (by simp [le_trans zero_le_one hq1] : 0 ≤ 1 / q), + refine lintegral_mono (λ a, ennreal.rpow_le_rpow _ (le_trans zero_le_one hq1)), + simp [←ennreal.coe_add, nnnorm_add_le], +end +... ≤ snorm' f q μ + snorm' g q μ : + ennreal.lintegral_Lp_add_le hf.ennnorm hg.ennnorm hq1 + +lemma snorm'_add_le_of_le_one {f g : α → E} + (hf : ae_strongly_measurable f μ) (hq0 : 0 ≤ q) (hq1 : q ≤ 1) : + snorm' (f + g) q μ ≤ 2^(1/q-1) * (snorm' f q μ + snorm' g q μ) := +calc (∫⁻ a, ↑‖(f + g) a‖₊ ^ q ∂μ) ^ (1 / q) + ≤ (∫⁻ a, (((λ a, (‖f a‖₊ : ℝ≥0∞)) + + (λ a, (‖g a‖₊ : ℝ≥0∞))) a) ^ q ∂μ) ^ (1 / q) : +begin + refine ennreal.rpow_le_rpow _ (by simp [hq0] : 0 ≤ 1 / q), + refine lintegral_mono (λ a, ennreal.rpow_le_rpow _ hq0), + simp [←ennreal.coe_add, nnnorm_add_le], +end +... ≤ 2^(1/q-1) * (snorm' f q μ + snorm' g q μ) : + ennreal.lintegral_Lp_add_le_of_le_one hf.ennnorm hq0 hq1 + +lemma snorm_ess_sup_add_le {f g : α → F} : + snorm_ess_sup (f + g) μ ≤ snorm_ess_sup f μ + snorm_ess_sup g μ := +begin + refine le_trans (ess_sup_mono_ae (eventually_of_forall (λ x, _))) + (ennreal.ess_sup_add_le _ _), + simp_rw [pi.add_apply, ←ennreal.coe_add, ennreal.coe_le_coe], + exact nnnorm_add_le _ _, +end + +lemma snorm_add_le + {f g : α → E} (hf : ae_strongly_measurable f μ) (hg : ae_strongly_measurable g μ) (hp1 : 1 ≤ p) : + snorm (f + g) p μ ≤ snorm f p μ + snorm g p μ := +begin + by_cases hp0 : p = 0, + { simp [hp0], }, + by_cases hp_top : p = ∞, + { simp [hp_top, snorm_ess_sup_add_le], }, + have hp1_real : 1 ≤ p.to_real, + by rwa [← ennreal.one_to_real, ennreal.to_real_le_to_real ennreal.one_ne_top hp_top], + repeat { rw snorm_eq_snorm' hp0 hp_top, }, + exact snorm'_add_le hf hg hp1_real, +end + +/-- A constant for the inequality `‖f + g‖_{L^p} ≤ C * (‖f‖_{L^p} + ‖g‖_{L^p})`. It is equal to `1` +for `p ≥ 1` or `p = 0`, and `2^(1/p-1)` in the more tricky interval `(0, 1)`. -/ +def Lp_add_const (p : ℝ≥0∞) : ℝ≥0∞ := +if p ∈ set.Ioo (0 : ℝ≥0∞) 1 then 2^(1/p.to_real-1) else 1 + +lemma Lp_add_const_of_one_le {p : ℝ≥0∞} (hp : 1 ≤ p) : Lp_add_const p = 1 := +begin + rw [Lp_add_const, if_neg], + assume h, + exact lt_irrefl _ (h.2.trans_le hp), +end + +lemma Lp_add_const_zero : Lp_add_const 0 = 1 := +begin + rw [Lp_add_const, if_neg], + assume h, + exact lt_irrefl _ (h.1), +end + +lemma Lp_add_const_lt_top (p : ℝ≥0∞) : Lp_add_const p < ∞ := +begin + rw [Lp_add_const], + split_ifs, + { apply ennreal.rpow_lt_top_of_nonneg _ ennreal.two_ne_top, + simp only [one_div, sub_nonneg], + apply one_le_inv (ennreal.to_real_pos h.1.ne' (h.2.trans ennreal.one_lt_top).ne), + simpa using ennreal.to_real_mono ennreal.one_ne_top h.2.le }, + { exact ennreal.one_lt_top } +end + +lemma snorm_add_le' + {f g : α → E} (hf : ae_strongly_measurable f μ) (hg : ae_strongly_measurable g μ) (p : ℝ≥0∞) : + snorm (f + g) p μ ≤ Lp_add_const p * (snorm f p μ + snorm g p μ) := +begin + rcases eq_or_ne p 0 with rfl|hp, + { simp only [snorm_exponent_zero, add_zero, mul_zero, le_zero_iff] }, + rcases lt_or_le p 1 with h'p|h'p, + { simp only [snorm_eq_snorm' hp (h'p.trans ennreal.one_lt_top).ne], + convert snorm'_add_le_of_le_one hf ennreal.to_real_nonneg _, + { have : p ∈ set.Ioo (0 : ℝ≥0∞) 1 := ⟨hp.bot_lt, h'p⟩, + simp only [Lp_add_const, if_pos this] }, + { simpa using ennreal.to_real_mono ennreal.one_ne_top h'p.le } }, + { simp [Lp_add_const_of_one_le h'p], + exact snorm_add_le hf hg h'p } +end + +variables (μ E) +/-- Technical lemma to control the addition of functions in `L^p` even for `p < 1`: Given `δ > 0`, +there exists `η` such that two functions bounded by `η` in `L^p` have a sum bounded by `δ`. One +could take `η = δ / 2` for `p ≥ 1`, but the point of the lemma is that it works also for `p < 1`. +-/ +lemma exists_Lp_half (p : ℝ≥0∞) {δ : ℝ≥0∞} (hδ : δ ≠ 0) : + ∃ (η : ℝ≥0∞), 0 < η ∧ ∀ (f g : α → E) (hf : ae_strongly_measurable f μ) + (hg : ae_strongly_measurable g μ) (Hf : snorm f p μ ≤ η) (Hg : snorm g p μ ≤ η), + snorm (f + g) p μ < δ := +begin + have : tendsto (λ (η : ℝ≥0∞), Lp_add_const p * (η + η)) (𝓝[>] 0) (𝓝 (Lp_add_const p * (0 + 0))), + from (ennreal.tendsto.const_mul (tendsto_id.add tendsto_id) + (or.inr (Lp_add_const_lt_top p).ne)).mono_left nhds_within_le_nhds, + simp only [add_zero, mul_zero] at this, + rcases (((tendsto_order.1 this).2 δ hδ.bot_lt).and self_mem_nhds_within).exists + with ⟨η, hη, ηpos⟩, + refine ⟨η, ηpos, λ f g hf hg Hf Hg, _⟩, + calc snorm (f + g) p μ ≤ Lp_add_const p * (snorm f p μ + snorm g p μ) : snorm_add_le' hf hg p + ... ≤ Lp_add_const p * (η + η) : mul_le_mul_of_nonneg_left (add_le_add Hf Hg) bot_le + ... < δ : hη +end +variables {μ E} + +lemma snorm_sub_le' + {f g : α → E} (hf : ae_strongly_measurable f μ) (hg : ae_strongly_measurable g μ) (p : ℝ≥0∞) : + snorm (f - g) p μ ≤ Lp_add_const p * (snorm f p μ + snorm g p μ) := +calc snorm (f - g) p μ = snorm (f + - g) p μ : by rw sub_eq_add_neg + -- We cannot use snorm_add_le on f and (-g) because we don't have `ae_measurable (-g) μ`, since + -- we don't suppose `[borel_space E]`. +... = snorm (λ x, ‖f x + - g x‖) p μ : (snorm_norm (f + - g)).symm +... ≤ snorm (λ x, ‖f x‖ + ‖- g x‖) p μ : by +{ refine snorm_mono_real (λ x, _), rw norm_norm, exact norm_add_le _ _, } +... = snorm (λ x, ‖f x‖ + ‖g x‖) p μ : by simp_rw norm_neg +... ≤ Lp_add_const p * (snorm (λ x, ‖f x‖) p μ + snorm (λ x, ‖g x‖) p μ) : + snorm_add_le' hf.norm hg.norm p +... = Lp_add_const p * (snorm f p μ + snorm g p μ) : by rw [← snorm_norm f, ← snorm_norm g] + +lemma snorm_sub_le + {f g : α → E} (hf : ae_strongly_measurable f μ) (hg : ae_strongly_measurable g μ) (hp : 1 ≤ p) : + snorm (f - g) p μ ≤ snorm f p μ + snorm g p μ := +by simpa [Lp_add_const_of_one_le hp] using snorm_sub_le' hf hg p + +lemma snorm_add_lt_top {f g : α → E} (hf : mem_ℒp f p μ) (hg : mem_ℒp g p μ) : + snorm (f + g) p μ < ∞ := calc +snorm (f + g) p μ ≤ Lp_add_const p * (snorm f p μ + snorm g p μ) : + snorm_add_le' hf.ae_strongly_measurable hg.ae_strongly_measurable p +... < ∞ : +begin + apply ennreal.mul_lt_top (Lp_add_const_lt_top p).ne, + exact ((ennreal.add_lt_top).2 ⟨hf.2, hg.2⟩).ne, +end + +lemma ae_le_snorm_ess_sup {f : α → F} : ∀ᵐ y ∂μ, ↑‖f y‖₊ ≤ snorm_ess_sup f μ := ae_le_ess_sup + +lemma meas_snorm_ess_sup_lt {f : α → F} : μ {y | snorm_ess_sup f μ < ‖f y‖₊} = 0 := +meas_ess_sup_lt + +section map_measure + +variables {β : Type*} {mβ : measurable_space β} {f : α → β} {g : β → E} + +include mβ + +lemma snorm_ess_sup_map_measure + (hg : ae_strongly_measurable g (measure.map f μ)) (hf : ae_measurable f μ) : + snorm_ess_sup g (measure.map f μ) = snorm_ess_sup (g ∘ f) μ := +ess_sup_map_measure hg.ennnorm hf + +lemma snorm_map_measure (hg : ae_strongly_measurable g (measure.map f μ)) (hf : ae_measurable f μ) : + snorm g p (measure.map f μ) = snorm (g ∘ f) p μ := +begin + by_cases hp_zero : p = 0, + { simp only [hp_zero, snorm_exponent_zero], }, + by_cases hp_top : p = ∞, + { simp_rw [hp_top, snorm_exponent_top], + exact snorm_ess_sup_map_measure hg hf, }, + simp_rw snorm_eq_lintegral_rpow_nnnorm hp_zero hp_top, + rw lintegral_map' (hg.ennnorm.pow_const p.to_real) hf, +end + +lemma mem_ℒp_map_measure_iff + (hg : ae_strongly_measurable g (measure.map f μ)) (hf : ae_measurable f μ) : + mem_ℒp g p (measure.map f μ) ↔ mem_ℒp (g ∘ f) p μ := +by simp [mem_ℒp, snorm_map_measure hg hf, hg.comp_ae_measurable hf, hg] + +lemma _root_.measurable_embedding.snorm_ess_sup_map_measure {g : β → F} + (hf : measurable_embedding f) : + snorm_ess_sup g (measure.map f μ) = snorm_ess_sup (g ∘ f) μ := +hf.ess_sup_map_measure + +lemma _root_.measurable_embedding.snorm_map_measure {g : β → F} (hf : measurable_embedding f) : + snorm g p (measure.map f μ) = snorm (g ∘ f) p μ := +begin + by_cases hp_zero : p = 0, + { simp only [hp_zero, snorm_exponent_zero], }, + by_cases hp : p = ∞, + { simp_rw [hp, snorm_exponent_top], + exact hf.ess_sup_map_measure, }, + { simp_rw snorm_eq_lintegral_rpow_nnnorm hp_zero hp, + rw hf.lintegral_map, }, +end + +lemma _root_.measurable_embedding.mem_ℒp_map_measure_iff {g : β → F} + (hf : measurable_embedding f) : + mem_ℒp g p (measure.map f μ) ↔ mem_ℒp (g ∘ f) p μ := +by simp_rw [mem_ℒp, hf.ae_strongly_measurable_map_iff, hf.snorm_map_measure] + +lemma _root_.measurable_equiv.mem_ℒp_map_measure_iff (f : α ≃ᵐ β) {g : β → F} : + mem_ℒp g p (measure.map f μ) ↔ mem_ℒp (g ∘ f) p μ := +f.measurable_embedding.mem_ℒp_map_measure_iff + +omit mβ + +end map_measure + +section trim + +lemma snorm'_trim (hm : m ≤ m0) {f : α → E} (hf : strongly_measurable[m] f) : + snorm' f q (ν.trim hm) = snorm' f q ν := +begin + simp_rw snorm', + congr' 1, + refine lintegral_trim hm _, + refine @measurable.pow_const _ _ _ _ _ _ _ m _ (@measurable.coe_nnreal_ennreal _ m _ _) _, + apply @strongly_measurable.measurable, + exact (@strongly_measurable.nnnorm α m _ _ _ hf), +end + +lemma limsup_trim (hm : m ≤ m0) {f : α → ℝ≥0∞} (hf : measurable[m] f) : + (ν.trim hm).ae.limsup f = ν.ae.limsup f := +begin + simp_rw limsup_eq, + suffices h_set_eq : {a : ℝ≥0∞ | ∀ᵐ n ∂(ν.trim hm), f n ≤ a} = {a : ℝ≥0∞ | ∀ᵐ n ∂ν, f n ≤ a}, + by rw h_set_eq, + ext1 a, + suffices h_meas_eq : ν {x | ¬ f x ≤ a} = ν.trim hm {x | ¬ f x ≤ a}, + by simp_rw [set.mem_set_of_eq, ae_iff, h_meas_eq], + refine (trim_measurable_set_eq hm _).symm, + refine @measurable_set.compl _ _ m (@measurable_set_le ℝ≥0∞ _ _ _ _ m _ _ _ _ _ hf _), + exact @measurable_const _ _ _ m _, +end + +lemma ess_sup_trim (hm : m ≤ m0) {f : α → ℝ≥0∞} (hf : measurable[m] f) : + ess_sup f (ν.trim hm) = ess_sup f ν := +by { simp_rw ess_sup, exact limsup_trim hm hf, } + +lemma snorm_ess_sup_trim (hm : m ≤ m0) {f : α → E} (hf : strongly_measurable[m] f) : + snorm_ess_sup f (ν.trim hm) = snorm_ess_sup f ν := +ess_sup_trim _ (@strongly_measurable.ennnorm _ m _ _ _ hf) + +lemma snorm_trim (hm : m ≤ m0) {f : α → E} (hf : strongly_measurable[m] f) : + snorm f p (ν.trim hm) = snorm f p ν := +begin + by_cases h0 : p = 0, + { simp [h0], }, + by_cases h_top : p = ∞, + { simpa only [h_top, snorm_exponent_top] using snorm_ess_sup_trim hm hf, }, + simpa only [snorm_eq_snorm' h0 h_top] using snorm'_trim hm hf, +end + +lemma snorm_trim_ae (hm : m ≤ m0) {f : α → E} (hf : ae_strongly_measurable f (ν.trim hm)) : + snorm f p (ν.trim hm) = snorm f p ν := +begin + rw [snorm_congr_ae hf.ae_eq_mk, snorm_congr_ae (ae_eq_of_ae_eq_trim hf.ae_eq_mk)], + exact snorm_trim hm hf.strongly_measurable_mk, +end + +lemma mem_ℒp_of_mem_ℒp_trim (hm : m ≤ m0) {f : α → E} (hf : mem_ℒp f p (ν.trim hm)) : + mem_ℒp f p ν := +⟨ae_strongly_measurable_of_ae_strongly_measurable_trim hm hf.1, +(le_of_eq (snorm_trim_ae hm hf.1).symm).trans_lt hf.2⟩ + +end trim + +@[simp] lemma snorm'_neg {f : α → F} : snorm' (-f) q μ = snorm' f q μ := by simp [snorm'] + +@[simp] lemma snorm_neg {f : α → F} : snorm (-f) p μ = snorm f p μ := +begin + by_cases h0 : p = 0, + { simp [h0], }, + by_cases h_top : p = ∞, + { simp [h_top, snorm_ess_sup], }, + simp [snorm_eq_snorm' h0 h_top], +end + +lemma mem_ℒp.neg {f : α → E} (hf : mem_ℒp f p μ) : mem_ℒp (-f) p μ := +⟨ae_strongly_measurable.neg hf.1, by simp [hf.right]⟩ + +lemma mem_ℒp_neg_iff {f : α → E} : mem_ℒp (-f) p μ ↔ mem_ℒp f p μ := +⟨λ h, neg_neg f ▸ h.neg, mem_ℒp.neg⟩ + +lemma snorm'_le_snorm'_mul_rpow_measure_univ {p q : ℝ} (hp0_lt : 0 < p) (hpq : p ≤ q) + {f : α → E} (hf : ae_strongly_measurable f μ) : + snorm' f p μ ≤ snorm' f q μ * (μ set.univ) ^ (1/p - 1/q) := +begin + have hq0_lt : 0 < q, from lt_of_lt_of_le hp0_lt hpq, + by_cases hpq_eq : p = q, + { rw [hpq_eq, sub_self, ennreal.rpow_zero, mul_one], + exact le_rfl, }, + have hpq : p < q, from lt_of_le_of_ne hpq hpq_eq, + let g := λ a : α, (1 : ℝ≥0∞), + have h_rw : ∫⁻ a, ↑‖f a‖₊^p ∂ μ = ∫⁻ a, (‖f a‖₊ * (g a))^p ∂ μ, + from lintegral_congr (λ a, by simp), + repeat {rw snorm'}, + rw h_rw, + let r := p * q / (q - p), + have hpqr : 1/p = 1/q + 1/r, + { field_simp [(ne_of_lt hp0_lt).symm, + (ne_of_lt hq0_lt).symm], + ring, }, + calc (∫⁻ (a : α), (↑‖f a‖₊ * g a) ^ p ∂μ) ^ (1/p) + ≤ (∫⁻ (a : α), ↑‖f a‖₊ ^ q ∂μ) ^ (1/q) * (∫⁻ (a : α), (g a) ^ r ∂μ) ^ (1/r) : + ennreal.lintegral_Lp_mul_le_Lq_mul_Lr hp0_lt hpq hpqr μ hf.ennnorm ae_measurable_const + ... = (∫⁻ (a : α), ↑‖f a‖₊ ^ q ∂μ) ^ (1/q) * μ set.univ ^ (1/p - 1/q) : + by simp [hpqr], +end + +lemma snorm'_le_snorm_ess_sup_mul_rpow_measure_univ (hq_pos : 0 < q) {f : α → F} : + snorm' f q μ ≤ snorm_ess_sup f μ * (μ set.univ) ^ (1/q) := +begin + have h_le : ∫⁻ (a : α), ↑‖f a‖₊ ^ q ∂μ ≤ ∫⁻ (a : α), (snorm_ess_sup f μ) ^ q ∂μ, + { refine lintegral_mono_ae _, + have h_nnnorm_le_snorm_ess_sup := coe_nnnorm_ae_le_snorm_ess_sup f μ, + refine h_nnnorm_le_snorm_ess_sup.mono (λ x hx, ennreal.rpow_le_rpow hx (le_of_lt hq_pos)), }, + rw [snorm', ←ennreal.rpow_one (snorm_ess_sup f μ)], + nth_rewrite 1 ←mul_inv_cancel (ne_of_lt hq_pos).symm, + rw [ennreal.rpow_mul, one_div, + ←ennreal.mul_rpow_of_nonneg _ _ (by simp [hq_pos.le] : 0 ≤ q⁻¹)], + refine ennreal.rpow_le_rpow _ (by simp [hq_pos.le]), + rwa lintegral_const at h_le, +end + +lemma snorm_le_snorm_mul_rpow_measure_univ {p q : ℝ≥0∞} (hpq : p ≤ q) {f : α → E} + (hf : ae_strongly_measurable f μ) : + snorm f p μ ≤ snorm f q μ * (μ set.univ) ^ (1/p.to_real - 1/q.to_real) := +begin + by_cases hp0 : p = 0, + { simp [hp0, zero_le], }, + rw ← ne.def at hp0, + have hp0_lt : 0 < p, from lt_of_le_of_ne (zero_le _) hp0.symm, + have hq0_lt : 0 < q, from lt_of_lt_of_le hp0_lt hpq, + by_cases hq_top : q = ∞, + { simp only [hq_top, div_zero, one_div, ennreal.top_to_real, sub_zero, snorm_exponent_top, + inv_zero], + by_cases hp_top : p = ∞, + { simp only [hp_top, ennreal.rpow_zero, mul_one, ennreal.top_to_real, sub_zero, inv_zero, + snorm_exponent_top], + exact le_rfl, }, + rw snorm_eq_snorm' hp0 hp_top, + have hp_pos : 0 < p.to_real, from ennreal.to_real_pos hp0_lt.ne' hp_top, + refine (snorm'_le_snorm_ess_sup_mul_rpow_measure_univ hp_pos).trans (le_of_eq _), + congr, + exact one_div _, }, + have hp_lt_top : p < ∞, from hpq.trans_lt (lt_top_iff_ne_top.mpr hq_top), + have hp_pos : 0 < p.to_real, from ennreal.to_real_pos hp0_lt.ne' hp_lt_top.ne, + rw [snorm_eq_snorm' hp0_lt.ne.symm hp_lt_top.ne, snorm_eq_snorm' hq0_lt.ne.symm hq_top], + have hpq_real : p.to_real ≤ q.to_real, by rwa ennreal.to_real_le_to_real hp_lt_top.ne hq_top, + exact snorm'_le_snorm'_mul_rpow_measure_univ hp_pos hpq_real hf, +end + +lemma snorm'_le_snorm'_of_exponent_le {m : measurable_space α} {p q : ℝ} (hp0_lt : 0 < p) + (hpq : p ≤ q) (μ : measure α) [is_probability_measure μ] {f : α → E} + (hf : ae_strongly_measurable f μ) : + snorm' f p μ ≤ snorm' f q μ := +begin + have h_le_μ := snorm'_le_snorm'_mul_rpow_measure_univ hp0_lt hpq hf, + rwa [measure_univ, ennreal.one_rpow, mul_one] at h_le_μ, +end + +lemma snorm'_le_snorm_ess_sup (hq_pos : 0 < q) {f : α → F} [is_probability_measure μ] : + snorm' f q μ ≤ snorm_ess_sup f μ := +le_trans (snorm'_le_snorm_ess_sup_mul_rpow_measure_univ hq_pos) (le_of_eq (by simp [measure_univ])) + +lemma snorm_le_snorm_of_exponent_le {p q : ℝ≥0∞} (hpq : p ≤ q) [is_probability_measure μ] + {f : α → E} (hf : ae_strongly_measurable f μ) : + snorm f p μ ≤ snorm f q μ := +(snorm_le_snorm_mul_rpow_measure_univ hpq hf).trans (le_of_eq (by simp [measure_univ])) + +lemma snorm'_lt_top_of_snorm'_lt_top_of_exponent_le {p q : ℝ} [is_finite_measure μ] {f : α → E} + (hf : ae_strongly_measurable f μ) (hfq_lt_top : snorm' f q μ < ∞) + (hp_nonneg : 0 ≤ p) (hpq : p ≤ q) : + snorm' f p μ < ∞ := +begin + cases le_or_lt p 0 with hp_nonpos hp_pos, + { rw le_antisymm hp_nonpos hp_nonneg, + simp, }, + have hq_pos : 0 < q, from lt_of_lt_of_le hp_pos hpq, + calc snorm' f p μ + ≤ snorm' f q μ * (μ set.univ) ^ (1/p - 1/q) : + snorm'_le_snorm'_mul_rpow_measure_univ hp_pos hpq hf + ... < ∞ : + begin + rw ennreal.mul_lt_top_iff, + refine or.inl ⟨hfq_lt_top, ennreal.rpow_lt_top_of_nonneg _ (measure_ne_top μ set.univ)⟩, + rwa [le_sub_comm, sub_zero, one_div, one_div, inv_le_inv hq_pos hp_pos], + end +end + +variables (μ) + +lemma pow_mul_meas_ge_le_snorm {f : α → E} + (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) (hf : ae_strongly_measurable f μ) (ε : ℝ≥0∞) : + (ε * μ {x | ε ≤ ‖f x‖₊ ^ p.to_real}) ^ (1 / p.to_real) ≤ snorm f p μ := +begin + rw snorm_eq_lintegral_rpow_nnnorm hp_ne_zero hp_ne_top, + exact ennreal.rpow_le_rpow (mul_meas_ge_le_lintegral₀ (hf.ennnorm.pow_const _) ε) + (one_div_nonneg.2 ennreal.to_real_nonneg), +end + +lemma mul_meas_ge_le_pow_snorm {f : α → E} + (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) (hf : ae_strongly_measurable f μ) (ε : ℝ≥0∞) : + ε * μ {x | ε ≤ ‖f x‖₊ ^ p.to_real} ≤ snorm f p μ ^ p.to_real := +begin + have : 1 / p.to_real * p.to_real = 1, + { refine one_div_mul_cancel _, + rw [ne, ennreal.to_real_eq_zero_iff], + exact not_or hp_ne_zero hp_ne_top }, + rw [← ennreal.rpow_one (ε * μ {x | ε ≤ ‖f x‖₊ ^ p.to_real}), ← this, ennreal.rpow_mul], + exact ennreal.rpow_le_rpow (pow_mul_meas_ge_le_snorm μ hp_ne_zero hp_ne_top hf ε) + ennreal.to_real_nonneg, +end + +/-- A version of Markov's inequality using Lp-norms. -/ +lemma mul_meas_ge_le_pow_snorm' {f : α → E} + (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) (hf : ae_strongly_measurable f μ) (ε : ℝ≥0∞) : + ε ^ p.to_real * μ {x | ε ≤ ‖f x‖₊} ≤ snorm f p μ ^ p.to_real := +begin + convert mul_meas_ge_le_pow_snorm μ hp_ne_zero hp_ne_top hf (ε ^ p.to_real), + ext x, + rw ennreal.rpow_le_rpow_iff (ennreal.to_real_pos hp_ne_zero hp_ne_top), +end + +lemma meas_ge_le_mul_pow_snorm {f : α → E} (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) + (hf : ae_strongly_measurable f μ) {ε : ℝ≥0∞} (hε : ε ≠ 0) : + μ {x | ε ≤ ‖f x‖₊} ≤ ε⁻¹ ^ p.to_real * snorm f p μ ^ p.to_real := +begin + by_cases ε = ∞, + { simp [h] }, + have hεpow : ε ^ p.to_real ≠ 0 := (ennreal.rpow_pos (pos_iff_ne_zero.2 hε) h).ne.symm, + have hεpow' : ε ^ p.to_real ≠ ∞ := (ennreal.rpow_ne_top_of_nonneg ennreal.to_real_nonneg h), + rw [ennreal.inv_rpow, ← ennreal.mul_le_mul_left hεpow hεpow', ← mul_assoc, + ennreal.mul_inv_cancel hεpow hεpow', one_mul], + exact mul_meas_ge_le_pow_snorm' μ hp_ne_zero hp_ne_top hf ε, +end + +variables {μ} + +lemma mem_ℒp.mem_ℒp_of_exponent_le {p q : ℝ≥0∞} [is_finite_measure μ] {f : α → E} + (hfq : mem_ℒp f q μ) (hpq : p ≤ q) : + mem_ℒp f p μ := +begin + cases hfq with hfq_m hfq_lt_top, + by_cases hp0 : p = 0, + { rwa [hp0, mem_ℒp_zero_iff_ae_strongly_measurable], }, + rw ←ne.def at hp0, + refine ⟨hfq_m, _⟩, + by_cases hp_top : p = ∞, + { have hq_top : q = ∞, + by rwa [hp_top, top_le_iff] at hpq, + rw [hp_top], + rwa hq_top at hfq_lt_top, }, + have hp_pos : 0 < p.to_real, from ennreal.to_real_pos hp0 hp_top, + by_cases hq_top : q = ∞, + { rw snorm_eq_snorm' hp0 hp_top, + rw [hq_top, snorm_exponent_top] at hfq_lt_top, + refine lt_of_le_of_lt (snorm'_le_snorm_ess_sup_mul_rpow_measure_univ hp_pos) _, + refine ennreal.mul_lt_top hfq_lt_top.ne _, + exact (ennreal.rpow_lt_top_of_nonneg (by simp [hp_pos.le]) (measure_ne_top μ set.univ)).ne }, + have hq0 : q ≠ 0, + { by_contra hq_eq_zero, + have hp_eq_zero : p = 0, from le_antisymm (by rwa hq_eq_zero at hpq) (zero_le _), + rw [hp_eq_zero, ennreal.zero_to_real] at hp_pos, + exact (lt_irrefl _) hp_pos, }, + have hpq_real : p.to_real ≤ q.to_real, by rwa ennreal.to_real_le_to_real hp_top hq_top, + rw snorm_eq_snorm' hp0 hp_top, + rw snorm_eq_snorm' hq0 hq_top at hfq_lt_top, + exact snorm'_lt_top_of_snorm'_lt_top_of_exponent_le hfq_m hfq_lt_top (le_of_lt hp_pos) hpq_real, +end + +section has_measurable_add +-- variable [has_measurable_add₂ E] + +lemma snorm'_sum_le {ι} {f : ι → α → E} {s : finset ι} + (hfs : ∀ i, i ∈ s → ae_strongly_measurable (f i) μ) (hq1 : 1 ≤ q) : + snorm' (∑ i in s, f i) q μ ≤ ∑ i in s, snorm' (f i) q μ := +finset.le_sum_of_subadditive_on_pred (λ (f : α → E), snorm' f q μ) + (λ f, ae_strongly_measurable f μ) (snorm'_zero (zero_lt_one.trans_le hq1)) + (λ f g hf hg, snorm'_add_le hf hg hq1) (λ f g hf hg, hf.add hg) _ hfs + +lemma snorm_sum_le {ι} {f : ι → α → E} {s : finset ι} + (hfs : ∀ i, i ∈ s → ae_strongly_measurable (f i) μ) (hp1 : 1 ≤ p) : + snorm (∑ i in s, f i) p μ ≤ ∑ i in s, snorm (f i) p μ := +finset.le_sum_of_subadditive_on_pred (λ (f : α → E), snorm f p μ) + (λ f, ae_strongly_measurable f μ) snorm_zero (λ f g hf hg, snorm_add_le hf hg hp1) + (λ f g hf hg, hf.add hg) _ hfs + +lemma mem_ℒp.add {f g : α → E} (hf : mem_ℒp f p μ) (hg : mem_ℒp g p μ) : mem_ℒp (f + g) p μ := +⟨ae_strongly_measurable.add hf.1 hg.1, snorm_add_lt_top hf hg⟩ + +lemma mem_ℒp.sub {f g : α → E} (hf : mem_ℒp f p μ) (hg : mem_ℒp g p μ) : mem_ℒp (f - g) p μ := +by { rw sub_eq_add_neg, exact hf.add hg.neg } + +lemma mem_ℒp_finset_sum {ι} (s : finset ι) {f : ι → α → E} (hf : ∀ i ∈ s, mem_ℒp (f i) p μ) : + mem_ℒp (λ a, ∑ i in s, f i a) p μ := +begin + haveI : decidable_eq ι := classical.dec_eq _, + revert hf, + refine finset.induction_on s _ _, + { simp only [zero_mem_ℒp', finset.sum_empty, implies_true_iff], }, + { intros i s his ih hf, + simp only [his, finset.sum_insert, not_false_iff], + 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 + +section monotonicity + +lemma snorm'_le_nnreal_smul_snorm'_of_ae_le_mul {f : α → F} {g : α → G} {c : ℝ≥0} + (h : ∀ᵐ x ∂μ, ‖f x‖₊ ≤ c * ‖g x‖₊) {p : ℝ} (hp : 0 < p) : + snorm' f p μ ≤ c • snorm' g p μ := +begin + simp_rw [snorm'], + rw [←ennreal.rpow_le_rpow_iff hp, ennreal.smul_def, smul_eq_mul, + ennreal.mul_rpow_of_nonneg _ _ hp.le], + simp_rw [←ennreal.rpow_mul, one_div, inv_mul_cancel hp.ne.symm, ennreal.rpow_one, + ennreal.coe_rpow_of_nonneg _ hp.le, ←lintegral_const_mul' _ _ ennreal.coe_ne_top, + ←ennreal.coe_mul], + apply lintegral_mono_ae, + simp_rw [ennreal.coe_le_coe, ←nnreal.mul_rpow, nnreal.rpow_le_rpow_iff hp], + exact h, +end + +lemma snorm_ess_sup_le_nnreal_smul_snorm_ess_sup_of_ae_le_mul {f : α → F} {g : α → G} {c : ℝ≥0} + (h : ∀ᵐ x ∂μ, ‖f x‖₊ ≤ c * ‖g x‖₊) : + snorm_ess_sup f μ ≤ c • snorm_ess_sup g μ := +calc ess_sup (λ x, (‖f x‖₊: ℝ≥0∞)) μ + ≤ ess_sup (λ x, (↑(c * ‖g x‖₊) : ℝ≥0∞)) μ + : ess_sup_mono_ae $ h.mono $ λ x hx, ennreal.coe_le_coe.mpr hx +... = ess_sup (λ x, (c * ‖g x‖₊ : ℝ≥0∞)) μ : by simp_rw ennreal.coe_mul +... = c • ess_sup (λ x, (‖g x‖₊ : ℝ≥0∞)) μ : ennreal.ess_sup_const_mul + +lemma snorm_le_nnreal_smul_snorm_of_ae_le_mul {f : α → F} {g : α → G} {c : ℝ≥0} + (h : ∀ᵐ x ∂μ, ‖f x‖₊ ≤ c * ‖g x‖₊) (p : ℝ≥0∞) : + snorm f p μ ≤ c • snorm g p μ := +begin + by_cases h0 : p = 0, + { simp [h0], }, + by_cases h_top : p = ∞, + { rw h_top, + exact snorm_ess_sup_le_nnreal_smul_snorm_ess_sup_of_ae_le_mul h, }, + simp_rw snorm_eq_snorm' h0 h_top, + exact snorm'_le_nnreal_smul_snorm'_of_ae_le_mul h (ennreal.to_real_pos h0 h_top), +end + +-- TODO: add the whole family of lemmas? +private lemma le_mul_iff_eq_zero_of_nonneg_of_neg_of_nonneg {α} [linear_ordered_semiring α] + {a b c : α} (ha : 0 ≤ a) (hb : b < 0) (hc : 0 ≤ c) : a ≤ b * c ↔ a = 0 ∧ c = 0 := +begin + split, + { intro h, + exact ⟨(h.trans (mul_nonpos_of_nonpos_of_nonneg hb.le hc)).antisymm ha, + (nonpos_of_mul_nonneg_right (ha.trans h) hb).antisymm hc⟩ }, + { rintro ⟨rfl, rfl⟩, + rw mul_zero, } +end + +/-- When `c` is negative, `‖f x‖ ≤ c * ‖g x‖` is nonsense and forces both `f` and `g` to have an +`snorm` of `0`. -/ +lemma snorm_eq_zero_and_zero_of_ae_le_mul_neg {f : α → F} {g : α → G} {c : ℝ} + (h : ∀ᵐ x ∂μ, ‖f x‖ ≤ c * ‖g x‖) (hc : c < 0) (p : ℝ≥0∞) : + snorm f p μ = 0 ∧ snorm g p μ = 0 := +begin + simp_rw [le_mul_iff_eq_zero_of_nonneg_of_neg_of_nonneg (norm_nonneg _) hc (norm_nonneg _), + norm_eq_zero, eventually_and] at h, + change f =ᵐ[μ] 0 ∧ g =ᵐ[μ] 0 at h, + simp [snorm_congr_ae h.1, snorm_congr_ae h.2], +end + +lemma snorm_le_mul_snorm_of_ae_le_mul {f : α → F} {g : α → G} {c : ℝ} + (h : ∀ᵐ x ∂μ, ‖f x‖ ≤ c * ‖g x‖) (p : ℝ≥0∞) : + snorm f p μ ≤ (ennreal.of_real c) * snorm g p μ := +snorm_le_nnreal_smul_snorm_of_ae_le_mul + (h.mono $ λ x hx, hx.trans $ mul_le_mul_of_nonneg_right c.le_coe_to_nnreal (norm_nonneg _)) _ + +lemma mem_ℒp.of_nnnorm_le_mul {f : α → E} {g : α → F} {c : ℝ≥0} + (hg : mem_ℒp g p μ) (hf : ae_strongly_measurable f μ) (hfg : ∀ᵐ x ∂μ, ‖f x‖₊ ≤ c * ‖g x‖₊) : + mem_ℒp f p μ := +⟨hf, (snorm_le_nnreal_smul_snorm_of_ae_le_mul hfg p).trans_lt $ + ennreal.mul_lt_top ennreal.coe_ne_top hg.snorm_ne_top⟩ + +lemma mem_ℒp.of_le_mul {f : α → E} {g : α → F} {c : ℝ} + (hg : mem_ℒp g p μ) (hf : ae_strongly_measurable f μ) (hfg : ∀ᵐ x ∂μ, ‖f x‖ ≤ c * ‖g x‖) : + mem_ℒp f p μ := +⟨hf, (snorm_le_mul_snorm_of_ae_le_mul hfg p).trans_lt $ + ennreal.mul_lt_top ennreal.of_real_ne_top hg.snorm_ne_top⟩ + +lemma snorm'_le_snorm'_mul_snorm' {p q r : ℝ} + {f : α → E} (hf : ae_strongly_measurable f μ) {g : α → F} (hg : ae_strongly_measurable g μ) + (b : E → F → G) (h : ∀ᵐ x ∂μ, ‖b (f x) (g x)‖₊ ≤ ‖f x‖₊ * ‖g x‖₊) + (hp0_lt : 0 < p) (hpq : p < q) (hpqr : 1/p = 1/q + 1/r) : + snorm' (λ x, b (f x) (g x)) p μ ≤ snorm' f q μ * snorm' g r μ := +begin + rw snorm', + calc (∫⁻ (a : α), ↑‖b (f a) (g a)‖₊ ^ p ∂μ) ^ (1 / p) + ≤ (∫⁻ (a : α), ↑(‖f a‖₊ * ‖g a‖₊) ^ p ∂μ) ^ (1 / p) : + (ennreal.rpow_le_rpow_iff $ one_div_pos.mpr (hp0_lt)).mpr $ + lintegral_mono_ae $ h.mono $ λ a ha, (ennreal.rpow_le_rpow_iff (hp0_lt)).mpr $ + ennreal.coe_le_coe.mpr $ ha + ... ≤ _ : _, + simp_rw [snorm', ennreal.coe_mul], + exact ennreal.lintegral_Lp_mul_le_Lq_mul_Lr hp0_lt hpq hpqr μ hf.ennnorm + hg.ennnorm, +end + +lemma snorm_le_snorm_top_mul_snorm (p : ℝ≥0∞) + (f : α → E) {g : α → F} (hg : ae_strongly_measurable g μ) (b : E → F → G) + (h : ∀ᵐ x ∂μ, ‖b (f x) (g x)‖₊ ≤ ‖f x‖₊ * ‖g x‖₊) : + snorm (λ x, b (f x) (g x)) p μ ≤ snorm f ∞ μ * snorm g p μ := +begin + by_cases hp_top : p = ∞, + { simp_rw [hp_top, snorm_exponent_top], + refine le_trans (ess_sup_mono_ae $ h.mono $ λ a ha, _) (ennreal.ess_sup_mul_le _ _), + simp_rw [pi.mul_apply, ←ennreal.coe_mul, ennreal.coe_le_coe], + exact ha }, + by_cases hp_zero : p = 0, + { simp only [hp_zero, snorm_exponent_zero, mul_zero, le_zero_iff], }, + simp_rw [snorm_eq_lintegral_rpow_nnnorm hp_zero hp_top, snorm_exponent_top, snorm_ess_sup], + calc (∫⁻ x, ↑‖b (f x) (g x)‖₊ ^ p.to_real ∂μ) ^ (1 / p.to_real) + ≤ (∫⁻ x, ↑‖f x‖₊ ^ p.to_real * ↑‖g x‖₊ ^ p.to_real ∂μ) ^ (1 / p.to_real) : + begin + refine ennreal.rpow_le_rpow _ (one_div_nonneg.mpr ennreal.to_real_nonneg), + refine lintegral_mono_ae (h.mono $ λ a ha, _), + rw ←ennreal.mul_rpow_of_nonneg _ _ ennreal.to_real_nonneg, + refine ennreal.rpow_le_rpow _ ennreal.to_real_nonneg, + rw [←ennreal.coe_mul, ennreal.coe_le_coe], + exact ha, + end + ... ≤ (∫⁻ x, (ess_sup (λ x, ↑‖f x‖₊) μ) ^ p.to_real * ↑‖g x‖₊ ^ p.to_real ∂μ) ^ (1 / p.to_real) : + begin + refine ennreal.rpow_le_rpow _ _, + swap, { rw one_div_nonneg, exact ennreal.to_real_nonneg, }, + refine lintegral_mono_ae _, + filter_upwards [@ennreal.ae_le_ess_sup _ _ μ (λ x, ↑‖f x‖₊)] with x hx, + exact mul_le_mul_right' (ennreal.rpow_le_rpow hx ennreal.to_real_nonneg) _ + end + ... = ess_sup (λ x, ↑‖f x‖₊) μ * (∫⁻ x, ↑‖g x‖₊ ^ p.to_real ∂μ) ^ (1 / p.to_real) : + begin + rw lintegral_const_mul'', + swap, { exact hg.nnnorm.ae_measurable.coe_nnreal_ennreal.pow ae_measurable_const, }, + rw ennreal.mul_rpow_of_nonneg, + swap, { rw one_div_nonneg, exact ennreal.to_real_nonneg, }, + rw [← ennreal.rpow_mul, one_div, mul_inv_cancel, ennreal.rpow_one], + rw [ne.def, ennreal.to_real_eq_zero_iff, auto.not_or_eq], + exact ⟨hp_zero, hp_top⟩, + end +end + +lemma snorm_le_snorm_mul_snorm_top (p : ℝ≥0∞) + {f : α → E} (hf : ae_strongly_measurable f μ) (g : α → F) (b : E → F → G) + (h : ∀ᵐ x ∂μ, ‖b (f x) (g x)‖₊ ≤ ‖f x‖₊ * ‖g x‖₊) : + snorm (λ x, b (f x) (g x)) p μ ≤ snorm f p μ * snorm g ∞ μ := +begin + rw [← snorm_norm f, ← snorm_norm g], + refine (snorm_mono_ae_real h).trans _, + simp_rw [mul_comm ‖f _‖₊, nnreal.coe_mul, coe_nnnorm], + rw mul_comm, + refine snorm_le_snorm_top_mul_snorm p (λ x, ‖g x‖) hf.norm _ (h.mono $ λ x hx, _), + simp_rw [nnnorm_mul], +end + +/-- Hölder's inequality, as an inequality on the `ℒp` seminorm of an elementwise operation +`λ x, b (f x) (g x)`. -/ +lemma snorm_le_snorm_mul_snorm_of_nnnorm {p q r : ℝ≥0∞} + {f : α → E} (hf : ae_strongly_measurable f μ) {g : α → F} (hg : ae_strongly_measurable g μ) + (b : E → F → G) (h : ∀ᵐ x ∂μ, ‖b (f x) (g x)‖₊ ≤ ‖f x‖₊ * ‖g x‖₊) + (hpqr : 1/p = 1/q + 1/r) : + snorm (λ x, b (f x) (g x)) p μ ≤ snorm f q μ * snorm g r μ := +begin + by_cases hp_zero : p = 0, + { simp [hp_zero], }, + have hq_ne_zero : q ≠ 0, + { intro hq_zero, + simp only [hq_zero, hp_zero, one_div, ennreal.inv_zero, top_add, + ennreal.inv_eq_top] at hpqr, + exact hpqr, }, + have hr_ne_zero : r ≠ 0, + { intro hr_zero, + simp only [hr_zero, hp_zero, one_div, ennreal.inv_zero, add_top, + ennreal.inv_eq_top] at hpqr, + exact hpqr, }, + by_cases hq_top : q = ∞, + { have hpr : p = r, + { simpa only [hq_top, one_div, ennreal.div_top, zero_add, inv_inj] using hpqr, }, + rw [← hpr, hq_top], + exact snorm_le_snorm_top_mul_snorm p f hg b h, }, + by_cases hr_top : r = ∞, + { have hpq : p = q, + { simpa only [hr_top, one_div, ennreal.div_top, add_zero, inv_inj] using hpqr, }, + rw [← hpq, hr_top], + exact snorm_le_snorm_mul_snorm_top p hf g b h, }, + have hpq : p < q, + { suffices : 1 / q < 1 / p, + { rwa [one_div, one_div, ennreal.inv_lt_inv] at this, }, + rw hpqr, + refine ennreal.lt_add_right _ _, + { simp only [hq_ne_zero, one_div, ne.def, ennreal.inv_eq_top, not_false_iff], }, + { simp only [hr_top, one_div, ne.def, ennreal.inv_eq_zero, not_false_iff], }, }, + rw [snorm_eq_snorm' hp_zero (hpq.trans_le le_top).ne, snorm_eq_snorm' hq_ne_zero hq_top, + snorm_eq_snorm' hr_ne_zero hr_top], + refine snorm'_le_snorm'_mul_snorm' hf hg _ h _ _ _, + { exact ennreal.to_real_pos hp_zero (hpq.trans_le le_top).ne, }, + { exact ennreal.to_real_strict_mono hq_top hpq, }, + rw [← ennreal.one_to_real, ← ennreal.to_real_div, ← ennreal.to_real_div, ← ennreal.to_real_div, + hpqr, ennreal.to_real_add], + { simp only [hq_ne_zero, one_div, ne.def, ennreal.inv_eq_top, not_false_iff], }, + { simp only [hr_ne_zero, one_div, ne.def, ennreal.inv_eq_top, not_false_iff], }, +end + +/-- Hölder's inequality, as an inequality on the `ℒp` seminorm of an elementwise operation +`λ x, b (f x) (g x)`. -/ +lemma snorm_le_snorm_mul_snorm'_of_norm {p q r : ℝ≥0∞} + {f : α → E} (hf : ae_strongly_measurable f μ) {g : α → F} (hg : ae_strongly_measurable g μ) + (b : E → F → G) (h : ∀ᵐ x ∂μ, ‖b (f x) (g x)‖ ≤ ‖f x‖ * ‖g x‖) + (hpqr : 1/p = 1/q + 1/r) : + snorm (λ x, b (f x) (g x)) p μ ≤ snorm f q μ * snorm g r μ := +snorm_le_snorm_mul_snorm_of_nnnorm hf hg b h hpqr + +end monotonicity + +/-! +### Bounded actions by normed rings + +In this section we show inequalities on the norm. +-/ +section has_bounded_smul + +variables {𝕜 : Type*} [normed_ring 𝕜] [mul_action_with_zero 𝕜 E] [mul_action_with_zero 𝕜 F] +variables [has_bounded_smul 𝕜 E] [has_bounded_smul 𝕜 F] + +lemma snorm'_const_smul_le (c : 𝕜) (f : α → F) (hq_pos : 0 < q) : + snorm' (c • f) q μ ≤ ‖c‖₊ • snorm' f q μ := +snorm'_le_nnreal_smul_snorm'_of_ae_le_mul (eventually_of_forall $ λ a, nnnorm_smul_le _ _) hq_pos + +lemma snorm_ess_sup_const_smul_le (c : 𝕜) (f : α → F) : + snorm_ess_sup (c • f) μ ≤ ‖c‖₊ • snorm_ess_sup f μ := +snorm_ess_sup_le_nnreal_smul_snorm_ess_sup_of_ae_le_mul + (eventually_of_forall $ λ a, nnnorm_smul_le _ _) + +lemma snorm_const_smul_le (c : 𝕜) (f : α → F) : + snorm (c • f) p μ ≤ ‖c‖₊ • snorm f p μ := +snorm_le_nnreal_smul_snorm_of_ae_le_mul (eventually_of_forall $ λ a, nnnorm_smul_le _ _) _ + +lemma mem_ℒp.const_smul {f : α → E} (hf : mem_ℒp f p μ) (c : 𝕜) : + mem_ℒp (c • f) p μ := +⟨ae_strongly_measurable.const_smul hf.1 c, + (snorm_const_smul_le c f).trans_lt (ennreal.mul_lt_top ennreal.coe_ne_top hf.2.ne)⟩ + +lemma mem_ℒp.const_mul {R} [normed_ring R] {f : α → R} (hf : mem_ℒp f p μ) (c : R) : + mem_ℒp (λ x, c * f x) p μ := +hf.const_smul c + +lemma snorm'_smul_le_mul_snorm' {p q r : ℝ} + {f : α → E} (hf : ae_strongly_measurable f μ) {φ : α → 𝕜} (hφ : ae_strongly_measurable φ μ) + (hp0_lt : 0 < p) (hpq : p < q) (hpqr : 1/p = 1/q + 1/r) : + snorm' (φ • f) p μ ≤ snorm' φ q μ * snorm' f r μ := +snorm'_le_snorm'_mul_snorm' hφ hf (•) + (eventually_of_forall $ λ a, nnnorm_smul_le _ _) hp0_lt hpq hpqr + +lemma snorm_smul_le_snorm_top_mul_snorm (p : ℝ≥0∞) + {f : α → E} (hf : ae_strongly_measurable f μ) (φ : α → 𝕜) : + snorm (φ • f) p μ ≤ snorm φ ∞ μ * snorm f p μ := +(snorm_le_snorm_top_mul_snorm p φ hf (•) (eventually_of_forall $ λ a, nnnorm_smul_le _ _) : _) + +lemma snorm_smul_le_snorm_mul_snorm_top (p : ℝ≥0∞) + (f : α → E) {φ : α → 𝕜} (hφ : ae_strongly_measurable φ μ) : + snorm (φ • f) p μ ≤ snorm φ p μ * snorm f ∞ μ := +(snorm_le_snorm_mul_snorm_top p hφ f (•) (eventually_of_forall $ λ a, nnnorm_smul_le _ _) : _) + +/-- Hölder's inequality, as an inequality on the `ℒp` seminorm of a scalar product `φ • f`. -/ +lemma snorm_smul_le_mul_snorm {p q r : ℝ≥0∞} + {f : α → E} (hf : ae_strongly_measurable f μ) {φ : α → 𝕜} (hφ : ae_strongly_measurable φ μ) + (hpqr : 1/p = 1/q + 1/r) : + snorm (φ • f) p μ ≤ snorm φ q μ * snorm f r μ := +(snorm_le_snorm_mul_snorm_of_nnnorm hφ hf (•) + (eventually_of_forall $ λ a, nnnorm_smul_le _ _) hpqr : _) + +lemma mem_ℒp.smul {p q r : ℝ≥0∞} {f : α → E} {φ : α → 𝕜} + (hf : mem_ℒp f r μ) (hφ : mem_ℒp φ q μ) (hpqr : 1/p = 1/q + 1/r) : + mem_ℒp (φ • f) p μ := +⟨hφ.1.smul hf.1, (snorm_smul_le_mul_snorm hf.1 hφ.1 hpqr).trans_lt + (ennreal.mul_lt_top hφ.snorm_ne_top hf.snorm_ne_top)⟩ + +lemma mem_ℒp.smul_of_top_right {p : ℝ≥0∞} {f : α → E} {φ : α → 𝕜} + (hf : mem_ℒp f p μ) (hφ : mem_ℒp φ ∞ μ) : + mem_ℒp (φ • f) p μ := +by { apply hf.smul hφ, simp only [ennreal.div_top, zero_add] } + +lemma mem_ℒp.smul_of_top_left {p : ℝ≥0∞} {f : α → E} {φ : α → 𝕜} + (hf : mem_ℒp f ∞ μ) (hφ : mem_ℒp φ p μ) : + mem_ℒp (φ • f) p μ := +by { apply hf.smul hφ, simp only [ennreal.div_top, add_zero] } + +end has_bounded_smul + +/-! +### Bounded actions by normed division rings + +The inequalities in the previous section are now tight. +-/ +section normed_space + +variables {𝕜 : Type*} [normed_division_ring 𝕜] [mul_action_with_zero 𝕜 E] [module 𝕜 F] +variables [has_bounded_smul 𝕜 E] [has_bounded_smul 𝕜 F] + +lemma snorm'_const_smul {f : α → F} (c : 𝕜) (hq_pos : 0 < q) : + snorm' (c • f) q μ = ‖c‖₊ • snorm' f q μ := +begin + obtain rfl | hc := eq_or_ne c 0, + { simp [snorm', hq_pos], }, + refine le_antisymm (snorm'_const_smul_le _ _ hq_pos) _, + have : snorm' _ q μ ≤ _:= snorm'_const_smul_le (c⁻¹) (c • f) hq_pos, + rwa [inv_smul_smul₀ hc, nnnorm_inv, ennreal.le_inv_smul_iff (nnnorm_ne_zero_iff.mpr hc)] at this, +end + +lemma snorm_ess_sup_const_smul (c : 𝕜) (f : α → F) : + snorm_ess_sup (c • f) μ = (‖c‖₊ : ℝ≥0∞) * snorm_ess_sup f μ := +by simp_rw [snorm_ess_sup, pi.smul_apply, nnnorm_smul, ennreal.coe_mul, ennreal.ess_sup_const_mul] + +lemma snorm_const_smul (c : 𝕜) (f : α → F) : + snorm (c • f) p μ = (‖c‖₊ : ℝ≥0∞) * snorm f p μ := +begin + obtain rfl | hc := eq_or_ne c 0, + { simp, }, + refine le_antisymm (snorm_const_smul_le _ _) _, + have : snorm _ p μ ≤ _:= snorm_const_smul_le (c⁻¹) (c • f), + rwa [inv_smul_smul₀ hc, nnnorm_inv, ennreal.le_inv_smul_iff (nnnorm_ne_zero_iff.mpr hc)] at this, +end + +end normed_space + +lemma snorm_indicator_ge_of_bdd_below (hp : p ≠ 0) (hp' : p ≠ ∞) + {f : α → F} (C : ℝ≥0) {s : set α} (hs : measurable_set s) + (hf : ∀ᵐ x ∂μ, x ∈ s → C ≤ ‖s.indicator f x‖₊) : + C • μ s ^ (1 / p.to_real) ≤ snorm (s.indicator f) p μ := +begin + rw [ennreal.smul_def, smul_eq_mul, snorm_eq_lintegral_rpow_nnnorm hp hp', + ennreal.le_rpow_one_div_iff (ennreal.to_real_pos hp hp'), + ennreal.mul_rpow_of_nonneg _ _ ennreal.to_real_nonneg, + ← ennreal.rpow_mul, one_div_mul_cancel (ennreal.to_real_pos hp hp').ne.symm, ennreal.rpow_one, + ← set_lintegral_const, ← lintegral_indicator _ hs], + refine lintegral_mono_ae _, + filter_upwards [hf] with x hx, + rw nnnorm_indicator_eq_indicator_nnnorm, + by_cases hxs : x ∈ s, + { simp only [set.indicator_of_mem hxs] at ⊢ hx, + exact ennreal.rpow_le_rpow (ennreal.coe_le_coe.2 (hx hxs)) ennreal.to_real_nonneg }, + { simp [set.indicator_of_not_mem hxs] }, +end + +section is_R_or_C +variables {𝕜 : Type*} [is_R_or_C 𝕜] {f : α → 𝕜} + +lemma mem_ℒp.re (hf : mem_ℒp f p μ) : mem_ℒp (λ x, is_R_or_C.re (f x)) p μ := +begin + have : ∀ x, ‖is_R_or_C.re (f x)‖ ≤ 1 * ‖f x‖, + by { intro x, rw one_mul, exact is_R_or_C.norm_re_le_norm (f x), }, + refine hf.of_le_mul _ (eventually_of_forall this), + exact is_R_or_C.continuous_re.comp_ae_strongly_measurable hf.1, +end + +lemma mem_ℒp.im (hf : mem_ℒp f p μ) : mem_ℒp (λ x, is_R_or_C.im (f x)) p μ := +begin + have : ∀ x, ‖is_R_or_C.im (f x)‖ ≤ 1 * ‖f x‖, + by { intro x, rw one_mul, exact is_R_or_C.norm_im_le_norm (f x), }, + refine hf.of_le_mul _ (eventually_of_forall this), + exact is_R_or_C.continuous_im.comp_ae_strongly_measurable hf.1, +end + +end is_R_or_C + +section liminf + +variables [measurable_space E] [opens_measurable_space E] {R : ℝ≥0} + +lemma ae_bdd_liminf_at_top_rpow_of_snorm_bdd {p : ℝ≥0∞} + {f : ℕ → α → E} (hfmeas : ∀ n, measurable (f n)) (hbdd : ∀ n, snorm (f n) p μ ≤ R) : + ∀ᵐ x ∂μ, liminf (λ n, (‖f n x‖₊ ^ p.to_real : ℝ≥0∞)) at_top < ∞ := +begin + by_cases hp0 : p.to_real = 0, + { simp only [hp0, ennreal.rpow_zero], + refine eventually_of_forall (λ x, _), + rw liminf_const (1 : ℝ≥0∞), + exacts [ennreal.one_lt_top, at_top_ne_bot] }, + have hp : p ≠ 0 := λ h, by simpa [h] using hp0, + have hp' : p ≠ ∞ := λ h, by simpa [h] using hp0, + refine ae_lt_top + (measurable_liminf (λ n, (hfmeas n).nnnorm.coe_nnreal_ennreal.pow_const p.to_real)) + (lt_of_le_of_lt (lintegral_liminf_le + (λ n, (hfmeas n).nnnorm.coe_nnreal_ennreal.pow_const p.to_real)) + (lt_of_le_of_lt _ (ennreal.rpow_lt_top_of_nonneg + ennreal.to_real_nonneg ennreal.coe_ne_top : ↑R ^ p.to_real < ∞))).ne, + simp_rw snorm_eq_lintegral_rpow_nnnorm hp hp' at hbdd, + simp_rw [liminf_eq, eventually_at_top], + exact Sup_le (λ b ⟨a, ha⟩, (ha a le_rfl).trans + ((ennreal.rpow_one_div_le_iff (ennreal.to_real_pos hp hp')).1 (hbdd _))), +end + +lemma ae_bdd_liminf_at_top_of_snorm_bdd {p : ℝ≥0∞} (hp : p ≠ 0) + {f : ℕ → α → E} (hfmeas : ∀ n, measurable (f n)) (hbdd : ∀ n, snorm (f n) p μ ≤ R) : + ∀ᵐ x ∂μ, liminf (λ n, (‖f n x‖₊ : ℝ≥0∞)) at_top < ∞ := +begin + by_cases hp' : p = ∞, + { subst hp', + simp_rw snorm_exponent_top at hbdd, + have : ∀ n, ∀ᵐ x ∂μ, (‖f n x‖₊ : ℝ≥0∞) < R + 1 := + λ n, ae_lt_of_ess_sup_lt (lt_of_le_of_lt (hbdd n) $ + ennreal.lt_add_right ennreal.coe_ne_top one_ne_zero), + rw ← ae_all_iff at this, + filter_upwards [this] with x hx using lt_of_le_of_lt + (liminf_le_of_frequently_le' $ frequently_of_forall $ λ n, (hx n).le) + (ennreal.add_lt_top.2 ⟨ennreal.coe_lt_top, ennreal.one_lt_top⟩) }, + filter_upwards [ae_bdd_liminf_at_top_rpow_of_snorm_bdd hfmeas hbdd] with x hx, + have hppos : 0 < p.to_real := ennreal.to_real_pos hp hp', + have : liminf (λ n, (‖f n x‖₊ ^ p.to_real : ℝ≥0∞)) at_top = + (liminf (λ n, (‖f n x‖₊ : ℝ≥0∞)) at_top)^ p.to_real, + { change liminf (λ n, ennreal.order_iso_rpow p.to_real hppos (‖f n x‖₊ : ℝ≥0∞)) at_top = + ennreal.order_iso_rpow p.to_real hppos (liminf (λ n, (‖f n x‖₊ : ℝ≥0∞)) at_top), + refine (order_iso.liminf_apply (ennreal.order_iso_rpow p.to_real _) _ _ _ _).symm; + is_bounded_default }, + rw this at hx, + rw [← ennreal.rpow_one (liminf (λ n, ‖f n x‖₊) at_top), ← mul_inv_cancel hppos.ne.symm, + ennreal.rpow_mul], + exact ennreal.rpow_lt_top_of_nonneg (inv_nonneg.2 hppos.le) hx.ne, +end + +end liminf + +end ℒp + +end measure_theory diff --git a/src/measure_theory/function/lp_space.lean b/src/measure_theory/function/lp_space.lean index 3617d299a125a..4e52fbbdf66e5 100644 --- a/src/measure_theory/function/lp_space.lean +++ b/src/measure_theory/function/lp_space.lean @@ -3,35 +3,19 @@ Copyright (c) 2020 Rémy Degenne. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne, Sébastien Gouëzel -/ -import analysis.normed_space.indicator_function import analysis.normed.group.hom -import analysis.special_functions.pow.continuity -import measure_theory.function.ess_sup -import measure_theory.function.ae_eq_fun -import measure_theory.integral.mean_inequalities +import measure_theory.function.lp_seminorm import topology.continuous_function.compact /-! -# ℒp space and Lp space +# Lp space -This file describes properties of almost everywhere strongly measurable functions with finite -seminorm, denoted by `snorm f p μ` and defined for `p:ℝ≥0∞` asmm_group (Lp E p μ) := `0` if `p=0`, -`(∫ ‖f a‖^p ∂μ) ^ (1/p)` for `0 < p < ∞` and `ess_sup ‖f‖ μ` for `p=∞`. - -The Prop-valued `mem_ℒp f p μ` states that a function `f : α → E` has finite seminorm. -The space `Lp E p μ` is the subtype of elements of `α →ₘ[μ] E` (see ae_eq_fun) such that -`snorm f p μ` is finite. For `1 ≤ p`, `snorm` defines a norm and `Lp` is a complete metric space. +This file provides the space `Lp E p μ` as the subtype of elements of `α →ₘ[μ] E` (see ae_eq_fun) +such that `snorm f p μ` is finite. For `1 ≤ p`, `snorm` defines a norm and `Lp` is a complete metric +space. ## Main definitions -* `snorm' f p μ` : `(∫ ‖f a‖^p ∂μ) ^ (1/p)` for `f : α → F` and `p : ℝ`, where `α` is a measurable - space and `F` is a normed group. -* `snorm_ess_sup f μ` : seminorm in `ℒ∞`, equal to the essential supremum `ess_sup ‖f‖ μ`. -* `snorm f p μ` : for `p : ℝ≥0∞`, seminorm in `ℒp`, equal to `0` for `p=0`, to `snorm' f p μ` - for `0 < p < ∞` and to `snorm_ess_sup f μ` for `p = ∞`. - -* `mem_ℒp f p μ` : property that the function `f` is almost everywhere strongly measurable and has - finite `p`-seminorm for the measure `μ` (`snorm f p μ < ∞`) * `Lp E p μ` : elements of `α →ₘ[μ] E` (see ae_eq_fun) such that `snorm f p μ` is finite. Defined as an `add_subgroup` of `α →ₘ[μ] E`. @@ -83,1592 +67,6 @@ variables {α E F G : Type*} {m m0 : measurable_space α} {p : ℝ≥0∞} {q : namespace measure_theory -section ℒp - -/-! -### ℒp seminorm - -We define the ℒp seminorm, denoted by `snorm f p μ`. For real `p`, it is given by an integral -formula (for which we use the notation `snorm' f p μ`), and for `p = ∞` it is the essential -supremum (for which we use the notation `snorm_ess_sup f μ`). - -We also define a predicate `mem_ℒp f p μ`, requesting that a function is almost everywhere -measurable and has finite `snorm f p μ`. - -This paragraph is devoted to the basic properties of these definitions. It is constructed as -follows: for a given property, we prove it for `snorm'` and `snorm_ess_sup` when it makes sense, -deduce it for `snorm`, and translate it in terms of `mem_ℒp`. --/ - -section ℒp_space_definition - -/-- `(∫ ‖f a‖^q ∂μ) ^ (1/q)`, which is a seminorm on the space of measurable functions for which -this quantity is finite -/ -def snorm' {m : measurable_space α} (f : α → F) (q : ℝ) (μ : measure α) : ℝ≥0∞ := -(∫⁻ a, ‖f a‖₊^q ∂μ) ^ (1/q) - -/-- seminorm for `ℒ∞`, equal to the essential supremum of `‖f‖`. -/ -def snorm_ess_sup {m : measurable_space α} (f : α → F) (μ : measure α) := -ess_sup (λ x, (‖f x‖₊ : ℝ≥0∞)) μ - -/-- `ℒp` seminorm, equal to `0` for `p=0`, to `(∫ ‖f a‖^p ∂μ) ^ (1/p)` for `0 < p < ∞` and to -`ess_sup ‖f‖ μ` for `p = ∞`. -/ -def snorm {m : measurable_space α} (f : α → F) (p : ℝ≥0∞) (μ : measure α) : ℝ≥0∞ := -if p = 0 then 0 else (if p = ∞ then snorm_ess_sup f μ else snorm' f (ennreal.to_real p) μ) - -lemma snorm_eq_snorm' (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) {f : α → F} : - snorm f p μ = snorm' f (ennreal.to_real p) μ := -by simp [snorm, hp_ne_zero, hp_ne_top] - -lemma snorm_eq_lintegral_rpow_nnnorm (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) {f : α → F} : - snorm f p μ = (∫⁻ x, ‖f x‖₊ ^ p.to_real ∂μ) ^ (1 / p.to_real) := -by rw [snorm_eq_snorm' hp_ne_zero hp_ne_top, snorm'] - -lemma snorm_one_eq_lintegral_nnnorm {f : α → F} : snorm f 1 μ = ∫⁻ x, ‖f x‖₊ ∂μ := -by simp_rw [snorm_eq_lintegral_rpow_nnnorm one_ne_zero ennreal.coe_ne_top, ennreal.one_to_real, - one_div_one, ennreal.rpow_one] - -@[simp] lemma snorm_exponent_top {f : α → F} : snorm f ∞ μ = snorm_ess_sup f μ := by simp [snorm] - -/-- 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 α . 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 μ) : - ae_strongly_measurable f μ := h.1 - -lemma lintegral_rpow_nnnorm_eq_rpow_snorm' {f : α → F} (hq0_lt : 0 < q) : - ∫⁻ a, ‖f a‖₊ ^ q ∂μ = (snorm' f q μ) ^ q := -begin - rw [snorm', ←ennreal.rpow_mul, one_div, inv_mul_cancel, ennreal.rpow_one], - exact (ne_of_lt hq0_lt).symm, -end - -end ℒp_space_definition - -section top - -lemma mem_ℒp.snorm_lt_top {f : α → E} (hfp : mem_ℒp f p μ) : snorm f p μ < ∞ := hfp.2 - -lemma mem_ℒp.snorm_ne_top {f : α → E} (hfp : mem_ℒp f p μ) : snorm f p μ ≠ ∞ := ne_of_lt (hfp.2) - -lemma lintegral_rpow_nnnorm_lt_top_of_snorm'_lt_top {f : α → F} (hq0_lt : 0 < q) - (hfq : snorm' f q μ < ∞) : - ∫⁻ a, ‖f a‖₊ ^ q ∂μ < ∞ := -begin - rw lintegral_rpow_nnnorm_eq_rpow_snorm' hq0_lt, - exact ennreal.rpow_lt_top_of_nonneg (le_of_lt hq0_lt) (ne_of_lt hfq), -end - -lemma lintegral_rpow_nnnorm_lt_top_of_snorm_lt_top {f : α → F} (hp_ne_zero : p ≠ 0) - (hp_ne_top : p ≠ ∞) (hfp : snorm f p μ < ∞) : - ∫⁻ a, ‖f a‖₊ ^ p.to_real ∂μ < ∞ := -begin - apply lintegral_rpow_nnnorm_lt_top_of_snorm'_lt_top, - { exact ennreal.to_real_pos hp_ne_zero hp_ne_top }, - { simpa [snorm_eq_snorm' hp_ne_zero hp_ne_top] using hfp } -end - -lemma snorm_lt_top_iff_lintegral_rpow_nnnorm_lt_top {f : α → F} (hp_ne_zero : p ≠ 0) - (hp_ne_top : p ≠ ∞) : - snorm f p μ < ∞ ↔ ∫⁻ a, ‖f a‖₊ ^ p.to_real ∂μ < ∞ := -⟨lintegral_rpow_nnnorm_lt_top_of_snorm_lt_top hp_ne_zero hp_ne_top, - begin - intros h, - have hp' := ennreal.to_real_pos hp_ne_zero hp_ne_top, - have : 0 < 1 / p.to_real := div_pos zero_lt_one hp', - simpa [snorm_eq_lintegral_rpow_nnnorm hp_ne_zero hp_ne_top] using - ennreal.rpow_lt_top_of_nonneg (le_of_lt this) (ne_of_lt h) - end⟩ - -end top - -section zero - -@[simp] lemma snorm'_exponent_zero {f : α → F} : snorm' f 0 μ = 1 := -by rw [snorm', div_zero, ennreal.rpow_zero] - -@[simp] lemma snorm_exponent_zero {f : α → F} : snorm f 0 μ = 0 := -by simp [snorm] - -lemma mem_ℒp_zero_iff_ae_strongly_measurable {f : α → E} : - mem_ℒp f 0 μ ↔ ae_strongly_measurable f μ := -by simp [mem_ℒp, snorm_exponent_zero] - -@[simp] lemma snorm'_zero (hp0_lt : 0 < q) : snorm' (0 : α → F) q μ = 0 := -by simp [snorm', hp0_lt] - -@[simp] lemma snorm'_zero' (hq0_ne : q ≠ 0) (hμ : μ ≠ 0) : snorm' (0 : α → F) q μ = 0 := -begin - cases le_or_lt 0 q with hq0 hq_neg, - { exact snorm'_zero (lt_of_le_of_ne hq0 hq0_ne.symm), }, - { simp [snorm', ennreal.rpow_eq_zero_iff, hμ, hq_neg], }, -end - -@[simp] lemma snorm_ess_sup_zero : snorm_ess_sup (0 : α → F) μ = 0 := -begin - simp_rw [snorm_ess_sup, pi.zero_apply, nnnorm_zero, ennreal.coe_zero, ←ennreal.bot_eq_zero], - exact ess_sup_const_bot, -end - -@[simp] lemma snorm_zero : snorm (0 : α → F) p μ = 0 := -begin - by_cases h0 : p = 0, - { simp [h0], }, - by_cases h_top : p = ∞, - { simp only [h_top, snorm_exponent_top, snorm_ess_sup_zero], }, - rw ←ne.def at h0, - simp [snorm_eq_snorm' h0 h_top, ennreal.to_real_pos h0 h_top], -end - -@[simp] lemma snorm_zero' : snorm (λ x : α, (0 : F)) p μ = 0 := -by convert snorm_zero - -lemma zero_mem_ℒp : mem_ℒp (0 : α → E) p μ := -⟨ae_strongly_measurable_zero, by { rw snorm_zero, exact ennreal.coe_lt_top, } ⟩ - -lemma zero_mem_ℒp' : mem_ℒp (λ x : α, (0 : E)) p μ := -by convert zero_mem_ℒp - -variables [measurable_space α] - -lemma snorm'_measure_zero_of_pos {f : α → F} (hq_pos : 0 < q) : - snorm' f q (0 : measure α) = 0 := -by simp [snorm', hq_pos] - -lemma snorm'_measure_zero_of_exponent_zero {f : α → F} : snorm' f 0 (0 : measure α) = 1 := -by simp [snorm'] - -lemma snorm'_measure_zero_of_neg {f : α → F} (hq_neg : q < 0) : snorm' f q (0 : measure α) = ∞ := -by simp [snorm', hq_neg] - -@[simp] lemma snorm_ess_sup_measure_zero {f : α → F} : snorm_ess_sup f (0 : measure α) = 0 := -by simp [snorm_ess_sup] - -@[simp] lemma snorm_measure_zero {f : α → F} : snorm f p (0 : measure α) = 0 := -begin - by_cases h0 : p = 0, - { simp [h0], }, - by_cases h_top : p = ∞, - { simp [h_top], }, - rw ←ne.def at h0, - simp [snorm_eq_snorm' h0 h_top, snorm', ennreal.to_real_pos h0 h_top], -end - -end zero - -section const - -lemma snorm'_const (c : F) (hq_pos : 0 < q) : - snorm' (λ x : α , c) q μ = (‖c‖₊ : ℝ≥0∞) * (μ set.univ) ^ (1/q) := -begin - rw [snorm', lintegral_const, ennreal.mul_rpow_of_nonneg _ _ (by simp [hq_pos.le] : 0 ≤ 1 / q)], - congr, - rw ←ennreal.rpow_mul, - suffices hq_cancel : q * (1/q) = 1, by rw [hq_cancel, ennreal.rpow_one], - rw [one_div, mul_inv_cancel (ne_of_lt hq_pos).symm], -end - -lemma snorm'_const' [is_finite_measure μ] (c : F) (hc_ne_zero : c ≠ 0) (hq_ne_zero : q ≠ 0) : - snorm' (λ x : α , c) q μ = (‖c‖₊ : ℝ≥0∞) * (μ set.univ) ^ (1/q) := -begin - rw [snorm', lintegral_const, ennreal.mul_rpow_of_ne_top _ (measure_ne_top μ set.univ)], - { congr, - rw ←ennreal.rpow_mul, - suffices hp_cancel : q * (1/q) = 1, by rw [hp_cancel, ennreal.rpow_one], - rw [one_div, mul_inv_cancel hq_ne_zero], }, - { rw [ne.def, ennreal.rpow_eq_top_iff, not_or_distrib, not_and_distrib, not_and_distrib], - split, - { left, - rwa [ennreal.coe_eq_zero, nnnorm_eq_zero], }, - { exact or.inl ennreal.coe_ne_top, }, }, -end - -lemma snorm_ess_sup_const (c : F) (hμ : μ ≠ 0) : - snorm_ess_sup (λ x : α, c) μ = (‖c‖₊ : ℝ≥0∞) := -by rw [snorm_ess_sup, ess_sup_const _ hμ] - -lemma snorm'_const_of_is_probability_measure (c : F) (hq_pos : 0 < q) [is_probability_measure μ] : - snorm' (λ x : α , c) q μ = (‖c‖₊ : ℝ≥0∞) := -by simp [snorm'_const c hq_pos, measure_univ] - -lemma snorm_const (c : F) (h0 : p ≠ 0) (hμ : μ ≠ 0) : - snorm (λ x : α , c) p μ = (‖c‖₊ : ℝ≥0∞) * (μ set.univ) ^ (1/(ennreal.to_real p)) := -begin - by_cases h_top : p = ∞, - { simp [h_top, snorm_ess_sup_const c hμ], }, - simp [snorm_eq_snorm' h0 h_top, snorm'_const, ennreal.to_real_pos h0 h_top], -end - -lemma snorm_const' (c : F) (h0 : p ≠ 0) (h_top: p ≠ ∞) : - snorm (λ x : α , c) p μ = (‖c‖₊ : ℝ≥0∞) * (μ set.univ) ^ (1/(ennreal.to_real p)) := -begin - simp [snorm_eq_snorm' h0 h_top, snorm'_const, ennreal.to_real_pos h0 h_top], -end - -lemma snorm_const_lt_top_iff {p : ℝ≥0∞} {c : F} (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) : - snorm (λ x : α, c) p μ < ∞ ↔ c = 0 ∨ μ set.univ < ∞ := -begin - have hp : 0 < p.to_real, from ennreal.to_real_pos hp_ne_zero hp_ne_top, - by_cases hμ : μ = 0, - { simp only [hμ, measure.coe_zero, pi.zero_apply, or_true, with_top.zero_lt_top, - snorm_measure_zero], }, - by_cases hc : c = 0, - { simp only [hc, true_or, eq_self_iff_true, with_top.zero_lt_top, snorm_zero'], }, - rw snorm_const' c hp_ne_zero hp_ne_top, - by_cases hμ_top : μ set.univ = ∞, - { simp [hc, hμ_top, hp], }, - rw ennreal.mul_lt_top_iff, - simp only [true_and, one_div, ennreal.rpow_eq_zero_iff, hμ, false_or, or_false, - ennreal.coe_lt_top, nnnorm_eq_zero, ennreal.coe_eq_zero, - measure_theory.measure.measure_univ_eq_zero, hp, inv_lt_zero, hc, and_false, false_and, - _root_.inv_pos, or_self, hμ_top, ne.lt_top hμ_top, iff_true], - exact ennreal.rpow_lt_top_of_nonneg (inv_nonneg.mpr hp.le) hμ_top, -end - -lemma mem_ℒp_const (c : E) [is_finite_measure μ] : mem_ℒp (λ a:α, c) p μ := -begin - refine ⟨ae_strongly_measurable_const, _⟩, - by_cases h0 : p = 0, - { simp [h0], }, - by_cases hμ : μ = 0, - { simp [hμ], }, - rw snorm_const c h0 hμ, - refine ennreal.mul_lt_top ennreal.coe_ne_top _, - refine (ennreal.rpow_lt_top_of_nonneg _ (measure_ne_top μ set.univ)).ne, - 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 - rw ← snorm_const_lt_top_iff hp_ne_zero hp_ne_top, - exact ⟨λ h, h.2, λ h, ⟨ae_strongly_measurable_const, h⟩⟩, -end - -end const - -lemma snorm'_mono_nnnorm_ae {f : α → F} {g : α → G} (hq : 0 ≤ q) (h : ∀ᵐ x ∂μ, ‖f x‖₊ ≤ ‖g x‖₊) : - snorm' f q μ ≤ snorm' g q μ := -begin - rw [snorm'], - refine ennreal.rpow_le_rpow _ (one_div_nonneg.2 hq), - refine lintegral_mono_ae (h.mono $ λ x hx, _), - exact ennreal.rpow_le_rpow (ennreal.coe_le_coe.2 hx) hq -end - -lemma snorm'_mono_ae {f : α → F} {g : α → G} (hq : 0 ≤ q) (h : ∀ᵐ x ∂μ, ‖f x‖ ≤ ‖g x‖) : - snorm' f q μ ≤ snorm' g q μ := -snorm'_mono_nnnorm_ae hq h - -lemma snorm'_congr_nnnorm_ae {f g : α → F} (hfg : ∀ᵐ x ∂μ, ‖f x‖₊ = ‖g x‖₊) : - snorm' f q μ = snorm' g q μ := -begin - have : (λ x, (‖f x‖₊ ^ q : ℝ≥0∞)) =ᵐ[μ] (λ x, ‖g x‖₊ ^ q), - from hfg.mono (λ x hx, by simp_rw hx), - simp only [snorm', lintegral_congr_ae this] -end - -lemma snorm'_congr_norm_ae {f g : α → F} (hfg : ∀ᵐ x ∂μ, ‖f x‖ = ‖g x‖) : - snorm' f q μ = snorm' g q μ := -snorm'_congr_nnnorm_ae $ hfg.mono $ λ x hx, nnreal.eq hx - -lemma snorm'_congr_ae {f g : α → F} (hfg : f =ᵐ[μ] g) : snorm' f q μ = snorm' g q μ := -snorm'_congr_nnnorm_ae (hfg.fun_comp _) - -lemma snorm_ess_sup_congr_ae {f g : α → F} (hfg : f =ᵐ[μ] g) : - snorm_ess_sup f μ = snorm_ess_sup g μ := -ess_sup_congr_ae (hfg.fun_comp (coe ∘ nnnorm)) - -lemma snorm_ess_sup_mono_nnnorm_ae {f g : α → F} (hfg : ∀ᵐ x ∂μ, ‖f x‖₊ ≤ ‖g x‖₊) : - snorm_ess_sup f μ ≤ snorm_ess_sup g μ := -ess_sup_mono_ae $ hfg.mono $ λ x hx, ennreal.coe_le_coe.mpr hx - -lemma snorm_mono_nnnorm_ae {f : α → F} {g : α → G} (h : ∀ᵐ x ∂μ, ‖f x‖₊ ≤ ‖g x‖₊) : - snorm f p μ ≤ snorm g p μ := -begin - simp only [snorm], - split_ifs, - { exact le_rfl }, - { exact ess_sup_mono_ae (h.mono $ λ x hx, ennreal.coe_le_coe.mpr hx) }, - { exact snorm'_mono_nnnorm_ae ennreal.to_real_nonneg h } -end - -lemma snorm_mono_ae {f : α → F} {g : α → G} (h : ∀ᵐ x ∂μ, ‖f x‖ ≤ ‖g x‖) : - snorm f p μ ≤ snorm g p μ := -snorm_mono_nnnorm_ae h - -lemma snorm_mono_ae_real {f : α → F} {g : α → ℝ} (h : ∀ᵐ x ∂μ, ‖f x‖ ≤ g x) : - snorm f p μ ≤ snorm g p μ := -snorm_mono_ae $ h.mono (λ x hx, hx.trans ((le_abs_self _).trans (real.norm_eq_abs _).symm.le)) - -lemma snorm_mono_nnnorm {f : α → F} {g : α → G} (h : ∀ x, ‖f x‖₊ ≤ ‖g x‖₊) : - snorm f p μ ≤ snorm g p μ := -snorm_mono_nnnorm_ae (eventually_of_forall (λ x, h x)) - -lemma snorm_mono {f : α → F} {g : α → G} (h : ∀ x, ‖f x‖ ≤ ‖g x‖) : - snorm f p μ ≤ snorm g p μ := -snorm_mono_nnnorm h - -lemma snorm_mono_real {f : α → F} {g : α → ℝ} (h : ∀ x, ‖f x‖ ≤ g x) : - snorm f p μ ≤ snorm g p μ := -snorm_mono_ae_real (eventually_of_forall (λ x, h x)) - -lemma snorm_ess_sup_le_of_ae_nnnorm_bound {f : α → F} {C : ℝ≥0} (hfC : ∀ᵐ x ∂μ, ‖f x‖₊ ≤ C) : - snorm_ess_sup f μ ≤ C := -ess_sup_le_of_ae_le C $ hfC.mono $ λ x hx, ennreal.coe_le_coe.mpr hx - -lemma snorm_ess_sup_le_of_ae_bound {f : α → F} {C : ℝ} (hfC : ∀ᵐ x ∂μ, ‖f x‖ ≤ C) : - snorm_ess_sup f μ ≤ ennreal.of_real C := -snorm_ess_sup_le_of_ae_nnnorm_bound $ hfC.mono $ λ x hx, hx.trans C.le_coe_to_nnreal - -lemma snorm_ess_sup_lt_top_of_ae_nnnorm_bound {f : α → F} {C : ℝ≥0} (hfC : ∀ᵐ x ∂μ, ‖f x‖₊ ≤ C) : - snorm_ess_sup f μ < ∞ := -(snorm_ess_sup_le_of_ae_nnnorm_bound hfC).trans_lt ennreal.coe_lt_top - -lemma snorm_ess_sup_lt_top_of_ae_bound {f : α → F} {C : ℝ} (hfC : ∀ᵐ x ∂μ, ‖f x‖ ≤ C) : - snorm_ess_sup f μ < ∞ := -(snorm_ess_sup_le_of_ae_bound hfC).trans_lt ennreal.of_real_lt_top - -lemma snorm_le_of_ae_nnnorm_bound {f : α → F} {C : ℝ≥0} (hfC : ∀ᵐ x ∂μ, ‖f x‖₊ ≤ C) : - snorm f p μ ≤ C • (μ set.univ ^ p.to_real⁻¹) := -begin - by_cases hμ : μ = 0, - { simp [hμ] }, - haveI : μ.ae.ne_bot := ae_ne_bot.mpr hμ, - by_cases hp : p = 0, - { simp [hp] }, - have : ∀ᵐ x ∂μ, ‖f x‖₊ ≤ ‖(C : ℝ)‖₊ := hfC.mono (λ x hx, hx.trans_eq C.nnnorm_eq.symm), - refine (snorm_mono_ae this).trans_eq _, - rw [snorm_const _ hp hμ, C.nnnorm_eq, one_div, ennreal.smul_def, smul_eq_mul], -end - -lemma snorm_le_of_ae_bound {f : α → F} {C : ℝ} (hfC : ∀ᵐ x ∂μ, ‖f x‖ ≤ C) : - snorm f p μ ≤ ((μ set.univ) ^ p.to_real⁻¹) * (ennreal.of_real C) := -begin - rw [←mul_comm], - exact snorm_le_of_ae_nnnorm_bound (hfC.mono $ λ x hx, hx.trans C.le_coe_to_nnreal), -end - -lemma snorm_congr_nnnorm_ae {f : α → F} {g : α → G} (hfg : ∀ᵐ x ∂μ, ‖f x‖₊ = ‖g x‖₊) : - snorm f p μ = snorm g p μ := -le_antisymm (snorm_mono_nnnorm_ae $ eventually_eq.le hfg) - (snorm_mono_nnnorm_ae $ (eventually_eq.symm hfg).le) - -lemma snorm_congr_norm_ae {f : α → F} {g : α → G} (hfg : ∀ᵐ x ∂μ, ‖f x‖ = ‖g x‖) : - snorm f p μ = snorm g p μ := -snorm_congr_nnnorm_ae $ hfg.mono $ λ x hx, nnreal.eq hx - -@[simp] lemma snorm'_norm {f : α → F} : snorm' (λ a, ‖f a‖) q μ = snorm' f q μ := -by simp [snorm'] - -@[simp] lemma snorm_norm (f : α → F) : snorm (λ x, ‖f x‖) p μ = snorm f p μ := -snorm_congr_norm_ae $ eventually_of_forall $ λ x, norm_norm _ - -lemma snorm'_norm_rpow (f : α → F) (p q : ℝ) (hq_pos : 0 < q) : - snorm' (λ x, ‖f x‖ ^ q) p μ = (snorm' f (p * q) μ) ^ q := -begin - simp_rw snorm', - rw [← ennreal.rpow_mul, ←one_div_mul_one_div], - simp_rw one_div, - rw [mul_assoc, inv_mul_cancel hq_pos.ne.symm, mul_one], - congr, - ext1 x, - simp_rw ← of_real_norm_eq_coe_nnnorm, - rw [real.norm_eq_abs, abs_eq_self.mpr (real.rpow_nonneg_of_nonneg (norm_nonneg _) _), - mul_comm, ← ennreal.of_real_rpow_of_nonneg (norm_nonneg _) hq_pos.le, ennreal.rpow_mul], -end - -lemma snorm_norm_rpow (f : α → F) (hq_pos : 0 < q) : - snorm (λ x, ‖f x‖ ^ q) p μ = (snorm f (p * ennreal.of_real q) μ) ^ q := -begin - by_cases h0 : p = 0, - { simp [h0, ennreal.zero_rpow_of_pos hq_pos], }, - by_cases hp_top : p = ∞, - { simp only [hp_top, snorm_exponent_top, ennreal.top_mul, hq_pos.not_le, ennreal.of_real_eq_zero, - if_false, snorm_exponent_top, snorm_ess_sup], - have h_rpow : ess_sup (λ (x : α), (‖(‖f x‖ ^ q)‖₊ : ℝ≥0∞)) μ - = ess_sup (λ (x : α), (↑‖f x‖₊) ^ q) μ, - { congr, - ext1 x, - nth_rewrite 1 ← nnnorm_norm, - rw [ennreal.coe_rpow_of_nonneg _ hq_pos.le, ennreal.coe_eq_coe], - ext, - push_cast, - rw real.norm_rpow_of_nonneg (norm_nonneg _), }, - rw h_rpow, - have h_rpow_mono := ennreal.strict_mono_rpow_of_pos hq_pos, - have h_rpow_surj := (ennreal.rpow_left_bijective hq_pos.ne.symm).2, - let iso := h_rpow_mono.order_iso_of_surjective _ h_rpow_surj, - exact (iso.ess_sup_apply (λ x, (‖f x‖₊ : ℝ≥0∞)) μ).symm, }, - rw [snorm_eq_snorm' h0 hp_top, snorm_eq_snorm' _ _], - swap, { refine mul_ne_zero h0 _, rwa [ne.def, ennreal.of_real_eq_zero, not_le], }, - swap, { exact ennreal.mul_ne_top hp_top ennreal.of_real_ne_top, }, - rw [ennreal.to_real_mul, ennreal.to_real_of_real hq_pos.le], - exact snorm'_norm_rpow f p.to_real q hq_pos, -end - -lemma snorm_congr_ae {f g : α → F} (hfg : f =ᵐ[μ] g) : snorm f p μ = snorm g p μ := -snorm_congr_norm_ae $ hfg.mono (λ x hx, hx ▸ rfl) - -lemma mem_ℒp_congr_ae {f g : α → E} (hfg : f =ᵐ[μ] g) : mem_ℒp f p μ ↔ mem_ℒp g p μ := -by simp only [mem_ℒp, snorm_congr_ae hfg, ae_strongly_measurable_congr hfg] - -lemma mem_ℒp.ae_eq {f g : α → E} (hfg : f =ᵐ[μ] g) (hf_Lp : mem_ℒp f p μ) : mem_ℒp g p μ := -(mem_ℒp_congr_ae hfg).1 hf_Lp - -lemma mem_ℒp.of_le {f : α → E} {g : α → F} - (hg : mem_ℒp g p μ) (hf : ae_strongly_measurable f μ) (hfg : ∀ᵐ x ∂μ, ‖f x‖ ≤ ‖g x‖) : - mem_ℒp f p μ := -⟨hf, (snorm_mono_ae hfg).trans_lt hg.snorm_lt_top⟩ - -alias mem_ℒp.of_le ← mem_ℒp.mono - -lemma mem_ℒp.mono' {f : α → E} {g : α → ℝ} (hg : mem_ℒp g p μ) - (hf : ae_strongly_measurable f μ) (h : ∀ᵐ a ∂μ, ‖f a‖ ≤ g a) : mem_ℒp f p μ := -hg.mono hf $ h.mono $ λ x hx, le_trans hx (le_abs_self _) - -lemma mem_ℒp.congr_norm {f : α → E} {g : α → F} (hf : mem_ℒp f p μ) - (hg : ae_strongly_measurable g μ) (h : ∀ᵐ a ∂μ, ‖f a‖ = ‖g a‖) : - mem_ℒp g p μ := -hf.mono hg $ eventually_eq.le $ eventually_eq.symm h - -lemma mem_ℒp_congr_norm {f : α → E} {g : α → F} - (hf : ae_strongly_measurable f μ) (hg : ae_strongly_measurable g μ) (h : ∀ᵐ a ∂μ, ‖f a‖ = ‖g a‖) : - mem_ℒp f p μ ↔ mem_ℒp g p μ := -⟨λ h2f, h2f.congr_norm hg h, λ h2g, h2g.congr_norm hf $ eventually_eq.symm h⟩ - -lemma mem_ℒp_top_of_bound {f : α → E} (hf : ae_strongly_measurable f μ) (C : ℝ) - (hfC : ∀ᵐ x ∂μ, ‖f x‖ ≤ C) : - mem_ℒp f ∞ μ := -⟨hf, by { rw snorm_exponent_top, exact snorm_ess_sup_lt_top_of_ae_bound hfC, }⟩ - -lemma mem_ℒp.of_bound [is_finite_measure μ] {f : α → E} (hf : ae_strongly_measurable f μ) - (C : ℝ) (hfC : ∀ᵐ x ∂μ, ‖f x‖ ≤ C) : - mem_ℒp f p μ := -(mem_ℒp_const C).of_le hf (hfC.mono (λ x hx, le_trans hx (le_abs_self _))) - -@[mono] lemma snorm'_mono_measure (f : α → F) (hμν : ν ≤ μ) (hq : 0 ≤ q) : - snorm' f q ν ≤ snorm' f q μ := -begin - simp_rw snorm', - suffices h_integral_mono : (∫⁻ a, (‖f a‖₊ : ℝ≥0∞) ^ q ∂ν) ≤ ∫⁻ a, ‖f a‖₊ ^ q ∂μ, - from ennreal.rpow_le_rpow h_integral_mono (by simp [hq]), - exact lintegral_mono' hμν le_rfl, -end - -@[mono] lemma snorm_ess_sup_mono_measure (f : α → F) (hμν : ν ≪ μ) : - snorm_ess_sup f ν ≤ snorm_ess_sup f μ := -by { simp_rw snorm_ess_sup, exact ess_sup_mono_measure hμν, } - -@[mono] lemma snorm_mono_measure (f : α → F) (hμν : ν ≤ μ) : - snorm f p ν ≤ snorm f p μ := -begin - by_cases hp0 : p = 0, - { simp [hp0], }, - by_cases hp_top : p = ∞, - { simp [hp_top, snorm_ess_sup_mono_measure f (measure.absolutely_continuous_of_le hμν)], }, - simp_rw snorm_eq_snorm' hp0 hp_top, - exact snorm'_mono_measure f hμν ennreal.to_real_nonneg, -end - -lemma mem_ℒp.mono_measure {f : α → E} (hμν : ν ≤ μ) (hf : mem_ℒp f p μ) : - mem_ℒp f p ν := -⟨hf.1.mono_measure hμν, (snorm_mono_measure f hμν).trans_lt hf.2⟩ - -lemma mem_ℒp.restrict (s : set α) {f : α → E} (hf : mem_ℒp f p μ) : - mem_ℒp f p (μ.restrict s) := -hf.mono_measure measure.restrict_le_self - -lemma snorm'_smul_measure {p : ℝ} (hp : 0 ≤ p) {f : α → F} (c : ℝ≥0∞) : - snorm' f p (c • μ) = c ^ (1 / p) * snorm' f p μ := -by { rw [snorm', lintegral_smul_measure, ennreal.mul_rpow_of_nonneg, snorm'], simp [hp], } - -lemma snorm_ess_sup_smul_measure {f : α → F} {c : ℝ≥0∞} (hc : c ≠ 0) : - snorm_ess_sup f (c • μ) = snorm_ess_sup f μ := -by { simp_rw [snorm_ess_sup], exact ess_sup_smul_measure hc, } - -/-- Use `snorm_smul_measure_of_ne_top` instead. -/ -private lemma snorm_smul_measure_of_ne_zero_of_ne_top {p : ℝ≥0∞} (hp_ne_zero : p ≠ 0) - (hp_ne_top : p ≠ ∞) {f : α → F} (c : ℝ≥0∞) : - snorm f p (c • μ) = c ^ (1 / p).to_real • snorm f p μ := -begin - simp_rw snorm_eq_snorm' hp_ne_zero hp_ne_top, - rw snorm'_smul_measure ennreal.to_real_nonneg, - congr, - simp_rw one_div, - rw ennreal.to_real_inv, -end - -lemma snorm_smul_measure_of_ne_zero {p : ℝ≥0∞} {f : α → F} {c : ℝ≥0∞} (hc : c ≠ 0) : - snorm f p (c • μ) = c ^ (1 / p).to_real • snorm f p μ := -begin - by_cases hp0 : p = 0, - { simp [hp0], }, - by_cases hp_top : p = ∞, - { simp [hp_top, snorm_ess_sup_smul_measure hc], }, - exact snorm_smul_measure_of_ne_zero_of_ne_top hp0 hp_top c, -end - -lemma snorm_smul_measure_of_ne_top {p : ℝ≥0∞} (hp_ne_top : p ≠ ∞) {f : α → F} (c : ℝ≥0∞) : - snorm f p (c • μ) = c ^ (1 / p).to_real • snorm f p μ := -begin - by_cases hp0 : p = 0, - { simp [hp0], }, - { exact snorm_smul_measure_of_ne_zero_of_ne_top hp0 hp_ne_top c, }, -end - -lemma snorm_one_smul_measure {f : α → F} (c : ℝ≥0∞) : - snorm f 1 (c • μ) = c * snorm f 1 μ := -by { rw @snorm_smul_measure_of_ne_top _ _ _ μ _ 1 (@ennreal.coe_ne_top 1) f c, simp, } - -lemma mem_ℒp.of_measure_le_smul {μ' : measure α} (c : ℝ≥0∞) (hc : c ≠ ∞) - (hμ'_le : μ' ≤ c • μ) {f : α → E} (hf : mem_ℒp f p μ) : - mem_ℒp f p μ' := -begin - refine ⟨hf.1.mono' (measure.absolutely_continuous_of_le_smul hμ'_le), _⟩, - refine (snorm_mono_measure f hμ'_le).trans_lt _, - by_cases hc0 : c = 0, - { simp [hc0], }, - rw [snorm_smul_measure_of_ne_zero hc0, smul_eq_mul], - refine ennreal.mul_lt_top _ hf.2.ne, - simp [hc, hc0], -end - -lemma mem_ℒp.smul_measure {f : α → E} {c : ℝ≥0∞} (hf : mem_ℒp f p μ) (hc : c ≠ ∞) : - mem_ℒp f p (c • μ) := -hf.of_measure_le_smul c hc le_rfl - -include m - -lemma snorm_one_add_measure (f : α → F) (μ ν : measure α) : - snorm f 1 (μ + ν) = snorm f 1 μ + snorm f 1 ν := -by { simp_rw snorm_one_eq_lintegral_nnnorm, rw lintegral_add_measure _ μ ν, } - -lemma snorm_le_add_measure_right (f : α → F) (μ ν : measure α) {p : ℝ≥0∞} : - snorm f p μ ≤ snorm f p (μ + ν) := -snorm_mono_measure f $ measure.le_add_right $ le_refl _ - -lemma snorm_le_add_measure_left (f : α → F) (μ ν : measure α) {p : ℝ≥0∞} : - snorm f p ν ≤ snorm f p (μ + ν) := -snorm_mono_measure f $ measure.le_add_left $ le_refl _ - -omit m - -lemma mem_ℒp.left_of_add_measure {f : α → E} (h : mem_ℒp f p (μ + ν)) : mem_ℒp f p μ := -h.mono_measure $ measure.le_add_right $ le_refl _ - -lemma mem_ℒp.right_of_add_measure {f : α → E} (h : mem_ℒp f p (μ + ν)) : mem_ℒp f p ν := -h.mono_measure $ measure.le_add_left $ le_refl _ - -lemma mem_ℒp.norm {f : α → E} (h : mem_ℒp f p μ) : mem_ℒp (λ x, ‖f x‖) p μ := -h.of_le h.ae_strongly_measurable.norm (eventually_of_forall (λ x, by simp)) - -lemma mem_ℒp_norm_iff {f : α → E} (hf : ae_strongly_measurable f μ) : - mem_ℒp (λ x, ‖f x‖) p μ ↔ mem_ℒp f p μ := -⟨λ h, ⟨hf, by { rw ← snorm_norm, exact h.2, }⟩, λ h, h.norm⟩ - -lemma snorm'_eq_zero_of_ae_zero {f : α → F} (hq0_lt : 0 < q) (hf_zero : f =ᵐ[μ] 0) : - snorm' f q μ = 0 := -by rw [snorm'_congr_ae hf_zero, snorm'_zero hq0_lt] - -lemma snorm'_eq_zero_of_ae_zero' (hq0_ne : q ≠ 0) (hμ : μ ≠ 0) {f : α → F} (hf_zero : f =ᵐ[μ] 0) : - snorm' f q μ = 0 := -by rw [snorm'_congr_ae hf_zero, snorm'_zero' hq0_ne hμ] - -lemma ae_eq_zero_of_snorm'_eq_zero {f : α → E} (hq0 : 0 ≤ q) (hf : ae_strongly_measurable f μ) - (h : snorm' f q μ = 0) : f =ᵐ[μ] 0 := -begin - rw [snorm', ennreal.rpow_eq_zero_iff] at h, - cases h, - { rw lintegral_eq_zero_iff' (hf.ennnorm.pow_const q) at h, - refine h.left.mono (λ x hx, _), - rw [pi.zero_apply, ennreal.rpow_eq_zero_iff] at hx, - cases hx, - { cases hx with hx _, - rwa [←ennreal.coe_zero, ennreal.coe_eq_coe, nnnorm_eq_zero] at hx, }, - { exact absurd hx.left ennreal.coe_ne_top, }, }, - { exfalso, - rw [one_div, inv_lt_zero] at h, - exact hq0.not_lt h.right }, -end - -lemma snorm'_eq_zero_iff (hq0_lt : 0 < q) {f : α → E} (hf : ae_strongly_measurable f μ) : - snorm' f q μ = 0 ↔ f =ᵐ[μ] 0 := -⟨ae_eq_zero_of_snorm'_eq_zero (le_of_lt hq0_lt) hf, snorm'_eq_zero_of_ae_zero hq0_lt⟩ - -lemma coe_nnnorm_ae_le_snorm_ess_sup {m : measurable_space α} (f : α → F) (μ : measure α) : - ∀ᵐ x ∂μ, (‖f x‖₊ : ℝ≥0∞) ≤ snorm_ess_sup f μ := -ennreal.ae_le_ess_sup (λ x, (‖f x‖₊ : ℝ≥0∞)) - -@[simp] lemma snorm_ess_sup_eq_zero_iff {f : α → F} : snorm_ess_sup f μ = 0 ↔ f =ᵐ[μ] 0 := -by simp [eventually_eq, snorm_ess_sup] - -lemma snorm_eq_zero_iff {f : α → E} (hf : ae_strongly_measurable f μ) (h0 : p ≠ 0) : - snorm f p μ = 0 ↔ f =ᵐ[μ] 0 := -begin - by_cases h_top : p = ∞, - { rw [h_top, snorm_exponent_top, snorm_ess_sup_eq_zero_iff], }, - rw snorm_eq_snorm' h0 h_top, - exact snorm'_eq_zero_iff (ennreal.to_real_pos h0 h_top) hf, -end - -lemma snorm'_add_le {f g : α → E} - (hf : ae_strongly_measurable f μ) (hg : ae_strongly_measurable g μ) (hq1 : 1 ≤ q) : - snorm' (f + g) q μ ≤ snorm' f q μ + snorm' g q μ := -calc (∫⁻ a, ↑‖(f + g) a‖₊ ^ q ∂μ) ^ (1 / q) - ≤ (∫⁻ a, (((λ a, (‖f a‖₊ : ℝ≥0∞)) - + (λ a, (‖g a‖₊ : ℝ≥0∞))) a) ^ q ∂μ) ^ (1 / q) : -begin - refine ennreal.rpow_le_rpow _ (by simp [le_trans zero_le_one hq1] : 0 ≤ 1 / q), - refine lintegral_mono (λ a, ennreal.rpow_le_rpow _ (le_trans zero_le_one hq1)), - simp [←ennreal.coe_add, nnnorm_add_le], -end -... ≤ snorm' f q μ + snorm' g q μ : - ennreal.lintegral_Lp_add_le hf.ennnorm hg.ennnorm hq1 - -lemma snorm'_add_le_of_le_one {f g : α → E} - (hf : ae_strongly_measurable f μ) (hq0 : 0 ≤ q) (hq1 : q ≤ 1) : - snorm' (f + g) q μ ≤ 2^(1/q-1) * (snorm' f q μ + snorm' g q μ) := -calc (∫⁻ a, ↑‖(f + g) a‖₊ ^ q ∂μ) ^ (1 / q) - ≤ (∫⁻ a, (((λ a, (‖f a‖₊ : ℝ≥0∞)) - + (λ a, (‖g a‖₊ : ℝ≥0∞))) a) ^ q ∂μ) ^ (1 / q) : -begin - refine ennreal.rpow_le_rpow _ (by simp [hq0] : 0 ≤ 1 / q), - refine lintegral_mono (λ a, ennreal.rpow_le_rpow _ hq0), - simp [←ennreal.coe_add, nnnorm_add_le], -end -... ≤ 2^(1/q-1) * (snorm' f q μ + snorm' g q μ) : - ennreal.lintegral_Lp_add_le_of_le_one hf.ennnorm hq0 hq1 - -lemma snorm_ess_sup_add_le {f g : α → F} : - snorm_ess_sup (f + g) μ ≤ snorm_ess_sup f μ + snorm_ess_sup g μ := -begin - refine le_trans (ess_sup_mono_ae (eventually_of_forall (λ x, _))) - (ennreal.ess_sup_add_le _ _), - simp_rw [pi.add_apply, ←ennreal.coe_add, ennreal.coe_le_coe], - exact nnnorm_add_le _ _, -end - -lemma snorm_add_le - {f g : α → E} (hf : ae_strongly_measurable f μ) (hg : ae_strongly_measurable g μ) (hp1 : 1 ≤ p) : - snorm (f + g) p μ ≤ snorm f p μ + snorm g p μ := -begin - by_cases hp0 : p = 0, - { simp [hp0], }, - by_cases hp_top : p = ∞, - { simp [hp_top, snorm_ess_sup_add_le], }, - have hp1_real : 1 ≤ p.to_real, - by rwa [← ennreal.one_to_real, ennreal.to_real_le_to_real ennreal.one_ne_top hp_top], - repeat { rw snorm_eq_snorm' hp0 hp_top, }, - exact snorm'_add_le hf hg hp1_real, -end - -/-- A constant for the inequality `‖f + g‖_{L^p} ≤ C * (‖f‖_{L^p} + ‖g‖_{L^p})`. It is equal to `1` -for `p ≥ 1` or `p = 0`, and `2^(1/p-1)` in the more tricky interval `(0, 1)`. -/ -def Lp_add_const (p : ℝ≥0∞) : ℝ≥0∞ := -if p ∈ set.Ioo (0 : ℝ≥0∞) 1 then 2^(1/p.to_real-1) else 1 - -lemma Lp_add_const_of_one_le {p : ℝ≥0∞} (hp : 1 ≤ p) : Lp_add_const p = 1 := -begin - rw [Lp_add_const, if_neg], - assume h, - exact lt_irrefl _ (h.2.trans_le hp), -end - -lemma Lp_add_const_zero : Lp_add_const 0 = 1 := -begin - rw [Lp_add_const, if_neg], - assume h, - exact lt_irrefl _ (h.1), -end - -lemma Lp_add_const_lt_top (p : ℝ≥0∞) : Lp_add_const p < ∞ := -begin - rw [Lp_add_const], - split_ifs, - { apply ennreal.rpow_lt_top_of_nonneg _ ennreal.two_ne_top, - simp only [one_div, sub_nonneg], - apply one_le_inv (ennreal.to_real_pos h.1.ne' (h.2.trans ennreal.one_lt_top).ne), - simpa using ennreal.to_real_mono ennreal.one_ne_top h.2.le }, - { exact ennreal.one_lt_top } -end - -lemma snorm_add_le' - {f g : α → E} (hf : ae_strongly_measurable f μ) (hg : ae_strongly_measurable g μ) (p : ℝ≥0∞) : - snorm (f + g) p μ ≤ Lp_add_const p * (snorm f p μ + snorm g p μ) := -begin - rcases eq_or_ne p 0 with rfl|hp, - { simp only [snorm_exponent_zero, add_zero, mul_zero, le_zero_iff] }, - rcases lt_or_le p 1 with h'p|h'p, - { simp only [snorm_eq_snorm' hp (h'p.trans ennreal.one_lt_top).ne], - convert snorm'_add_le_of_le_one hf ennreal.to_real_nonneg _, - { have : p ∈ set.Ioo (0 : ℝ≥0∞) 1 := ⟨hp.bot_lt, h'p⟩, - simp only [Lp_add_const, if_pos this] }, - { simpa using ennreal.to_real_mono ennreal.one_ne_top h'p.le } }, - { simp [Lp_add_const_of_one_le h'p], - exact snorm_add_le hf hg h'p } -end - -variables (μ E) -/-- Technical lemma to control the addition of functions in `L^p` even for `p < 1`: Given `δ > 0`, -there exists `η` such that two functions bounded by `η` in `L^p` have a sum bounded by `δ`. One -could take `η = δ / 2` for `p ≥ 1`, but the point of the lemma is that it works also for `p < 1`. --/ -lemma exists_Lp_half (p : ℝ≥0∞) {δ : ℝ≥0∞} (hδ : δ ≠ 0) : - ∃ (η : ℝ≥0∞), 0 < η ∧ ∀ (f g : α → E) (hf : ae_strongly_measurable f μ) - (hg : ae_strongly_measurable g μ) (Hf : snorm f p μ ≤ η) (Hg : snorm g p μ ≤ η), - snorm (f + g) p μ < δ := -begin - have : tendsto (λ (η : ℝ≥0∞), Lp_add_const p * (η + η)) (𝓝[>] 0) (𝓝 (Lp_add_const p * (0 + 0))), - from (ennreal.tendsto.const_mul (tendsto_id.add tendsto_id) - (or.inr (Lp_add_const_lt_top p).ne)).mono_left nhds_within_le_nhds, - simp only [add_zero, mul_zero] at this, - rcases (((tendsto_order.1 this).2 δ hδ.bot_lt).and self_mem_nhds_within).exists - with ⟨η, hη, ηpos⟩, - refine ⟨η, ηpos, λ f g hf hg Hf Hg, _⟩, - calc snorm (f + g) p μ ≤ Lp_add_const p * (snorm f p μ + snorm g p μ) : snorm_add_le' hf hg p - ... ≤ Lp_add_const p * (η + η) : mul_le_mul_of_nonneg_left (add_le_add Hf Hg) bot_le - ... < δ : hη -end -variables {μ E} - -lemma snorm_sub_le' - {f g : α → E} (hf : ae_strongly_measurable f μ) (hg : ae_strongly_measurable g μ) (p : ℝ≥0∞) : - snorm (f - g) p μ ≤ Lp_add_const p * (snorm f p μ + snorm g p μ) := -calc snorm (f - g) p μ = snorm (f + - g) p μ : by rw sub_eq_add_neg - -- We cannot use snorm_add_le on f and (-g) because we don't have `ae_measurable (-g) μ`, since - -- we don't suppose `[borel_space E]`. -... = snorm (λ x, ‖f x + - g x‖) p μ : (snorm_norm (f + - g)).symm -... ≤ snorm (λ x, ‖f x‖ + ‖- g x‖) p μ : by -{ refine snorm_mono_real (λ x, _), rw norm_norm, exact norm_add_le _ _, } -... = snorm (λ x, ‖f x‖ + ‖g x‖) p μ : by simp_rw norm_neg -... ≤ Lp_add_const p * (snorm (λ x, ‖f x‖) p μ + snorm (λ x, ‖g x‖) p μ) : - snorm_add_le' hf.norm hg.norm p -... = Lp_add_const p * (snorm f p μ + snorm g p μ) : by rw [← snorm_norm f, ← snorm_norm g] - -lemma snorm_sub_le - {f g : α → E} (hf : ae_strongly_measurable f μ) (hg : ae_strongly_measurable g μ) (hp : 1 ≤ p) : - snorm (f - g) p μ ≤ snorm f p μ + snorm g p μ := -by simpa [Lp_add_const_of_one_le hp] using snorm_sub_le' hf hg p - -lemma snorm_add_lt_top {f g : α → E} (hf : mem_ℒp f p μ) (hg : mem_ℒp g p μ) : - snorm (f + g) p μ < ∞ := calc -snorm (f + g) p μ ≤ Lp_add_const p * (snorm f p μ + snorm g p μ) : - snorm_add_le' hf.ae_strongly_measurable hg.ae_strongly_measurable p -... < ∞ : -begin - apply ennreal.mul_lt_top (Lp_add_const_lt_top p).ne, - exact ((ennreal.add_lt_top).2 ⟨hf.2, hg.2⟩).ne, -end - -lemma ae_le_snorm_ess_sup {f : α → F} : ∀ᵐ y ∂μ, ↑‖f y‖₊ ≤ snorm_ess_sup f μ := ae_le_ess_sup - -lemma meas_snorm_ess_sup_lt {f : α → F} : μ {y | snorm_ess_sup f μ < ‖f y‖₊} = 0 := -meas_ess_sup_lt - -section map_measure - -variables {β : Type*} {mβ : measurable_space β} {f : α → β} {g : β → E} - -include mβ - -lemma snorm_ess_sup_map_measure - (hg : ae_strongly_measurable g (measure.map f μ)) (hf : ae_measurable f μ) : - snorm_ess_sup g (measure.map f μ) = snorm_ess_sup (g ∘ f) μ := -ess_sup_map_measure hg.ennnorm hf - -lemma snorm_map_measure (hg : ae_strongly_measurable g (measure.map f μ)) (hf : ae_measurable f μ) : - snorm g p (measure.map f μ) = snorm (g ∘ f) p μ := -begin - by_cases hp_zero : p = 0, - { simp only [hp_zero, snorm_exponent_zero], }, - by_cases hp_top : p = ∞, - { simp_rw [hp_top, snorm_exponent_top], - exact snorm_ess_sup_map_measure hg hf, }, - simp_rw snorm_eq_lintegral_rpow_nnnorm hp_zero hp_top, - rw lintegral_map' (hg.ennnorm.pow_const p.to_real) hf, -end - -lemma mem_ℒp_map_measure_iff - (hg : ae_strongly_measurable g (measure.map f μ)) (hf : ae_measurable f μ) : - mem_ℒp g p (measure.map f μ) ↔ mem_ℒp (g ∘ f) p μ := -by simp [mem_ℒp, snorm_map_measure hg hf, hg.comp_ae_measurable hf, hg] - -lemma _root_.measurable_embedding.snorm_ess_sup_map_measure {g : β → F} - (hf : measurable_embedding f) : - snorm_ess_sup g (measure.map f μ) = snorm_ess_sup (g ∘ f) μ := -hf.ess_sup_map_measure - -lemma _root_.measurable_embedding.snorm_map_measure {g : β → F} (hf : measurable_embedding f) : - snorm g p (measure.map f μ) = snorm (g ∘ f) p μ := -begin - by_cases hp_zero : p = 0, - { simp only [hp_zero, snorm_exponent_zero], }, - by_cases hp : p = ∞, - { simp_rw [hp, snorm_exponent_top], - exact hf.ess_sup_map_measure, }, - { simp_rw snorm_eq_lintegral_rpow_nnnorm hp_zero hp, - rw hf.lintegral_map, }, -end - -lemma _root_.measurable_embedding.mem_ℒp_map_measure_iff {g : β → F} - (hf : measurable_embedding f) : - mem_ℒp g p (measure.map f μ) ↔ mem_ℒp (g ∘ f) p μ := -by simp_rw [mem_ℒp, hf.ae_strongly_measurable_map_iff, hf.snorm_map_measure] - -lemma _root_.measurable_equiv.mem_ℒp_map_measure_iff (f : α ≃ᵐ β) {g : β → F} : - mem_ℒp g p (measure.map f μ) ↔ mem_ℒp (g ∘ f) p μ := -f.measurable_embedding.mem_ℒp_map_measure_iff - -omit mβ - -end map_measure - -section trim - -lemma snorm'_trim (hm : m ≤ m0) {f : α → E} (hf : strongly_measurable[m] f) : - snorm' f q (ν.trim hm) = snorm' f q ν := -begin - simp_rw snorm', - congr' 1, - refine lintegral_trim hm _, - refine @measurable.pow_const _ _ _ _ _ _ _ m _ (@measurable.coe_nnreal_ennreal _ m _ _) _, - apply @strongly_measurable.measurable, - exact (@strongly_measurable.nnnorm α m _ _ _ hf), -end - -lemma limsup_trim (hm : m ≤ m0) {f : α → ℝ≥0∞} (hf : measurable[m] f) : - (ν.trim hm).ae.limsup f = ν.ae.limsup f := -begin - simp_rw limsup_eq, - suffices h_set_eq : {a : ℝ≥0∞ | ∀ᵐ n ∂(ν.trim hm), f n ≤ a} = {a : ℝ≥0∞ | ∀ᵐ n ∂ν, f n ≤ a}, - by rw h_set_eq, - ext1 a, - suffices h_meas_eq : ν {x | ¬ f x ≤ a} = ν.trim hm {x | ¬ f x ≤ a}, - by simp_rw [set.mem_set_of_eq, ae_iff, h_meas_eq], - refine (trim_measurable_set_eq hm _).symm, - refine @measurable_set.compl _ _ m (@measurable_set_le ℝ≥0∞ _ _ _ _ m _ _ _ _ _ hf _), - exact @measurable_const _ _ _ m _, -end - -lemma ess_sup_trim (hm : m ≤ m0) {f : α → ℝ≥0∞} (hf : measurable[m] f) : - ess_sup f (ν.trim hm) = ess_sup f ν := -by { simp_rw ess_sup, exact limsup_trim hm hf, } - -lemma snorm_ess_sup_trim (hm : m ≤ m0) {f : α → E} (hf : strongly_measurable[m] f) : - snorm_ess_sup f (ν.trim hm) = snorm_ess_sup f ν := -ess_sup_trim _ (@strongly_measurable.ennnorm _ m _ _ _ hf) - -lemma snorm_trim (hm : m ≤ m0) {f : α → E} (hf : strongly_measurable[m] f) : - snorm f p (ν.trim hm) = snorm f p ν := -begin - by_cases h0 : p = 0, - { simp [h0], }, - by_cases h_top : p = ∞, - { simpa only [h_top, snorm_exponent_top] using snorm_ess_sup_trim hm hf, }, - simpa only [snorm_eq_snorm' h0 h_top] using snorm'_trim hm hf, -end - -lemma snorm_trim_ae (hm : m ≤ m0) {f : α → E} (hf : ae_strongly_measurable f (ν.trim hm)) : - snorm f p (ν.trim hm) = snorm f p ν := -begin - rw [snorm_congr_ae hf.ae_eq_mk, snorm_congr_ae (ae_eq_of_ae_eq_trim hf.ae_eq_mk)], - exact snorm_trim hm hf.strongly_measurable_mk, -end - -lemma mem_ℒp_of_mem_ℒp_trim (hm : m ≤ m0) {f : α → E} (hf : mem_ℒp f p (ν.trim hm)) : - mem_ℒp f p ν := -⟨ae_strongly_measurable_of_ae_strongly_measurable_trim hm hf.1, -(le_of_eq (snorm_trim_ae hm hf.1).symm).trans_lt hf.2⟩ - -end trim - -@[simp] lemma snorm'_neg {f : α → F} : snorm' (-f) q μ = snorm' f q μ := by simp [snorm'] - -@[simp] lemma snorm_neg {f : α → F} : snorm (-f) p μ = snorm f p μ := -begin - by_cases h0 : p = 0, - { simp [h0], }, - by_cases h_top : p = ∞, - { simp [h_top, snorm_ess_sup], }, - simp [snorm_eq_snorm' h0 h_top], -end - -lemma mem_ℒp.neg {f : α → E} (hf : mem_ℒp f p μ) : mem_ℒp (-f) p μ := -⟨ae_strongly_measurable.neg hf.1, by simp [hf.right]⟩ - -lemma mem_ℒp_neg_iff {f : α → E} : mem_ℒp (-f) p μ ↔ mem_ℒp f p μ := -⟨λ h, neg_neg f ▸ h.neg, mem_ℒp.neg⟩ - -lemma snorm'_le_snorm'_mul_rpow_measure_univ {p q : ℝ} (hp0_lt : 0 < p) (hpq : p ≤ q) - {f : α → E} (hf : ae_strongly_measurable f μ) : - snorm' f p μ ≤ snorm' f q μ * (μ set.univ) ^ (1/p - 1/q) := -begin - have hq0_lt : 0 < q, from lt_of_lt_of_le hp0_lt hpq, - by_cases hpq_eq : p = q, - { rw [hpq_eq, sub_self, ennreal.rpow_zero, mul_one], - exact le_rfl, }, - have hpq : p < q, from lt_of_le_of_ne hpq hpq_eq, - let g := λ a : α, (1 : ℝ≥0∞), - have h_rw : ∫⁻ a, ↑‖f a‖₊^p ∂ μ = ∫⁻ a, (‖f a‖₊ * (g a))^p ∂ μ, - from lintegral_congr (λ a, by simp), - repeat {rw snorm'}, - rw h_rw, - let r := p * q / (q - p), - have hpqr : 1/p = 1/q + 1/r, - { field_simp [(ne_of_lt hp0_lt).symm, - (ne_of_lt hq0_lt).symm], - ring, }, - calc (∫⁻ (a : α), (↑‖f a‖₊ * g a) ^ p ∂μ) ^ (1/p) - ≤ (∫⁻ (a : α), ↑‖f a‖₊ ^ q ∂μ) ^ (1/q) * (∫⁻ (a : α), (g a) ^ r ∂μ) ^ (1/r) : - ennreal.lintegral_Lp_mul_le_Lq_mul_Lr hp0_lt hpq hpqr μ hf.ennnorm ae_measurable_const - ... = (∫⁻ (a : α), ↑‖f a‖₊ ^ q ∂μ) ^ (1/q) * μ set.univ ^ (1/p - 1/q) : - by simp [hpqr], -end - -lemma snorm'_le_snorm_ess_sup_mul_rpow_measure_univ (hq_pos : 0 < q) {f : α → F} : - snorm' f q μ ≤ snorm_ess_sup f μ * (μ set.univ) ^ (1/q) := -begin - have h_le : ∫⁻ (a : α), ↑‖f a‖₊ ^ q ∂μ ≤ ∫⁻ (a : α), (snorm_ess_sup f μ) ^ q ∂μ, - { refine lintegral_mono_ae _, - have h_nnnorm_le_snorm_ess_sup := coe_nnnorm_ae_le_snorm_ess_sup f μ, - refine h_nnnorm_le_snorm_ess_sup.mono (λ x hx, ennreal.rpow_le_rpow hx (le_of_lt hq_pos)), }, - rw [snorm', ←ennreal.rpow_one (snorm_ess_sup f μ)], - nth_rewrite 1 ←mul_inv_cancel (ne_of_lt hq_pos).symm, - rw [ennreal.rpow_mul, one_div, - ←ennreal.mul_rpow_of_nonneg _ _ (by simp [hq_pos.le] : 0 ≤ q⁻¹)], - refine ennreal.rpow_le_rpow _ (by simp [hq_pos.le]), - rwa lintegral_const at h_le, -end - -lemma snorm_le_snorm_mul_rpow_measure_univ {p q : ℝ≥0∞} (hpq : p ≤ q) {f : α → E} - (hf : ae_strongly_measurable f μ) : - snorm f p μ ≤ snorm f q μ * (μ set.univ) ^ (1/p.to_real - 1/q.to_real) := -begin - by_cases hp0 : p = 0, - { simp [hp0, zero_le], }, - rw ← ne.def at hp0, - have hp0_lt : 0 < p, from lt_of_le_of_ne (zero_le _) hp0.symm, - have hq0_lt : 0 < q, from lt_of_lt_of_le hp0_lt hpq, - by_cases hq_top : q = ∞, - { simp only [hq_top, div_zero, one_div, ennreal.top_to_real, sub_zero, snorm_exponent_top, - inv_zero], - by_cases hp_top : p = ∞, - { simp only [hp_top, ennreal.rpow_zero, mul_one, ennreal.top_to_real, sub_zero, inv_zero, - snorm_exponent_top], - exact le_rfl, }, - rw snorm_eq_snorm' hp0 hp_top, - have hp_pos : 0 < p.to_real, from ennreal.to_real_pos hp0_lt.ne' hp_top, - refine (snorm'_le_snorm_ess_sup_mul_rpow_measure_univ hp_pos).trans (le_of_eq _), - congr, - exact one_div _, }, - have hp_lt_top : p < ∞, from hpq.trans_lt (lt_top_iff_ne_top.mpr hq_top), - have hp_pos : 0 < p.to_real, from ennreal.to_real_pos hp0_lt.ne' hp_lt_top.ne, - rw [snorm_eq_snorm' hp0_lt.ne.symm hp_lt_top.ne, snorm_eq_snorm' hq0_lt.ne.symm hq_top], - have hpq_real : p.to_real ≤ q.to_real, by rwa ennreal.to_real_le_to_real hp_lt_top.ne hq_top, - exact snorm'_le_snorm'_mul_rpow_measure_univ hp_pos hpq_real hf, -end - -lemma snorm'_le_snorm'_of_exponent_le {m : measurable_space α} {p q : ℝ} (hp0_lt : 0 < p) - (hpq : p ≤ q) (μ : measure α) [is_probability_measure μ] {f : α → E} - (hf : ae_strongly_measurable f μ) : - snorm' f p μ ≤ snorm' f q μ := -begin - have h_le_μ := snorm'_le_snorm'_mul_rpow_measure_univ hp0_lt hpq hf, - rwa [measure_univ, ennreal.one_rpow, mul_one] at h_le_μ, -end - -lemma snorm'_le_snorm_ess_sup (hq_pos : 0 < q) {f : α → F} [is_probability_measure μ] : - snorm' f q μ ≤ snorm_ess_sup f μ := -le_trans (snorm'_le_snorm_ess_sup_mul_rpow_measure_univ hq_pos) (le_of_eq (by simp [measure_univ])) - -lemma snorm_le_snorm_of_exponent_le {p q : ℝ≥0∞} (hpq : p ≤ q) [is_probability_measure μ] - {f : α → E} (hf : ae_strongly_measurable f μ) : - snorm f p μ ≤ snorm f q μ := -(snorm_le_snorm_mul_rpow_measure_univ hpq hf).trans (le_of_eq (by simp [measure_univ])) - -lemma snorm'_lt_top_of_snorm'_lt_top_of_exponent_le {p q : ℝ} [is_finite_measure μ] {f : α → E} - (hf : ae_strongly_measurable f μ) (hfq_lt_top : snorm' f q μ < ∞) - (hp_nonneg : 0 ≤ p) (hpq : p ≤ q) : - snorm' f p μ < ∞ := -begin - cases le_or_lt p 0 with hp_nonpos hp_pos, - { rw le_antisymm hp_nonpos hp_nonneg, - simp, }, - have hq_pos : 0 < q, from lt_of_lt_of_le hp_pos hpq, - calc snorm' f p μ - ≤ snorm' f q μ * (μ set.univ) ^ (1/p - 1/q) : - snorm'_le_snorm'_mul_rpow_measure_univ hp_pos hpq hf - ... < ∞ : - begin - rw ennreal.mul_lt_top_iff, - refine or.inl ⟨hfq_lt_top, ennreal.rpow_lt_top_of_nonneg _ (measure_ne_top μ set.univ)⟩, - rwa [le_sub_comm, sub_zero, one_div, one_div, inv_le_inv hq_pos hp_pos], - end -end - -variables (μ) - -lemma pow_mul_meas_ge_le_snorm {f : α → E} - (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) (hf : ae_strongly_measurable f μ) (ε : ℝ≥0∞) : - (ε * μ {x | ε ≤ ‖f x‖₊ ^ p.to_real}) ^ (1 / p.to_real) ≤ snorm f p μ := -begin - rw snorm_eq_lintegral_rpow_nnnorm hp_ne_zero hp_ne_top, - exact ennreal.rpow_le_rpow (mul_meas_ge_le_lintegral₀ (hf.ennnorm.pow_const _) ε) - (one_div_nonneg.2 ennreal.to_real_nonneg), -end - -lemma mul_meas_ge_le_pow_snorm {f : α → E} - (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) (hf : ae_strongly_measurable f μ) (ε : ℝ≥0∞) : - ε * μ {x | ε ≤ ‖f x‖₊ ^ p.to_real} ≤ snorm f p μ ^ p.to_real := -begin - have : 1 / p.to_real * p.to_real = 1, - { refine one_div_mul_cancel _, - rw [ne, ennreal.to_real_eq_zero_iff], - exact not_or hp_ne_zero hp_ne_top }, - rw [← ennreal.rpow_one (ε * μ {x | ε ≤ ‖f x‖₊ ^ p.to_real}), ← this, ennreal.rpow_mul], - exact ennreal.rpow_le_rpow (pow_mul_meas_ge_le_snorm μ hp_ne_zero hp_ne_top hf ε) - ennreal.to_real_nonneg, -end - -/-- A version of Markov's inequality using Lp-norms. -/ -lemma mul_meas_ge_le_pow_snorm' {f : α → E} - (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) (hf : ae_strongly_measurable f μ) (ε : ℝ≥0∞) : - ε ^ p.to_real * μ {x | ε ≤ ‖f x‖₊} ≤ snorm f p μ ^ p.to_real := -begin - convert mul_meas_ge_le_pow_snorm μ hp_ne_zero hp_ne_top hf (ε ^ p.to_real), - ext x, - rw ennreal.rpow_le_rpow_iff (ennreal.to_real_pos hp_ne_zero hp_ne_top), -end - -lemma meas_ge_le_mul_pow_snorm {f : α → E} (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) - (hf : ae_strongly_measurable f μ) {ε : ℝ≥0∞} (hε : ε ≠ 0) : - μ {x | ε ≤ ‖f x‖₊} ≤ ε⁻¹ ^ p.to_real * snorm f p μ ^ p.to_real := -begin - by_cases ε = ∞, - { simp [h] }, - have hεpow : ε ^ p.to_real ≠ 0 := (ennreal.rpow_pos (pos_iff_ne_zero.2 hε) h).ne.symm, - have hεpow' : ε ^ p.to_real ≠ ∞ := (ennreal.rpow_ne_top_of_nonneg ennreal.to_real_nonneg h), - rw [ennreal.inv_rpow, ← ennreal.mul_le_mul_left hεpow hεpow', ← mul_assoc, - ennreal.mul_inv_cancel hεpow hεpow', one_mul], - exact mul_meas_ge_le_pow_snorm' μ hp_ne_zero hp_ne_top hf ε, -end - -variables {μ} - -lemma mem_ℒp.mem_ℒp_of_exponent_le {p q : ℝ≥0∞} [is_finite_measure μ] {f : α → E} - (hfq : mem_ℒp f q μ) (hpq : p ≤ q) : - mem_ℒp f p μ := -begin - cases hfq with hfq_m hfq_lt_top, - by_cases hp0 : p = 0, - { rwa [hp0, mem_ℒp_zero_iff_ae_strongly_measurable], }, - rw ←ne.def at hp0, - refine ⟨hfq_m, _⟩, - by_cases hp_top : p = ∞, - { have hq_top : q = ∞, - by rwa [hp_top, top_le_iff] at hpq, - rw [hp_top], - rwa hq_top at hfq_lt_top, }, - have hp_pos : 0 < p.to_real, from ennreal.to_real_pos hp0 hp_top, - by_cases hq_top : q = ∞, - { rw snorm_eq_snorm' hp0 hp_top, - rw [hq_top, snorm_exponent_top] at hfq_lt_top, - refine lt_of_le_of_lt (snorm'_le_snorm_ess_sup_mul_rpow_measure_univ hp_pos) _, - refine ennreal.mul_lt_top hfq_lt_top.ne _, - exact (ennreal.rpow_lt_top_of_nonneg (by simp [hp_pos.le]) (measure_ne_top μ set.univ)).ne }, - have hq0 : q ≠ 0, - { by_contra hq_eq_zero, - have hp_eq_zero : p = 0, from le_antisymm (by rwa hq_eq_zero at hpq) (zero_le _), - rw [hp_eq_zero, ennreal.zero_to_real] at hp_pos, - exact (lt_irrefl _) hp_pos, }, - have hpq_real : p.to_real ≤ q.to_real, by rwa ennreal.to_real_le_to_real hp_top hq_top, - rw snorm_eq_snorm' hp0 hp_top, - rw snorm_eq_snorm' hq0 hq_top at hfq_lt_top, - exact snorm'_lt_top_of_snorm'_lt_top_of_exponent_le hfq_m hfq_lt_top (le_of_lt hp_pos) hpq_real, -end - -section has_measurable_add --- variable [has_measurable_add₂ E] - -lemma snorm'_sum_le {ι} {f : ι → α → E} {s : finset ι} - (hfs : ∀ i, i ∈ s → ae_strongly_measurable (f i) μ) (hq1 : 1 ≤ q) : - snorm' (∑ i in s, f i) q μ ≤ ∑ i in s, snorm' (f i) q μ := -finset.le_sum_of_subadditive_on_pred (λ (f : α → E), snorm' f q μ) - (λ f, ae_strongly_measurable f μ) (snorm'_zero (zero_lt_one.trans_le hq1)) - (λ f g hf hg, snorm'_add_le hf hg hq1) (λ f g hf hg, hf.add hg) _ hfs - -lemma snorm_sum_le {ι} {f : ι → α → E} {s : finset ι} - (hfs : ∀ i, i ∈ s → ae_strongly_measurable (f i) μ) (hp1 : 1 ≤ p) : - snorm (∑ i in s, f i) p μ ≤ ∑ i in s, snorm (f i) p μ := -finset.le_sum_of_subadditive_on_pred (λ (f : α → E), snorm f p μ) - (λ f, ae_strongly_measurable f μ) snorm_zero (λ f g hf hg, snorm_add_le hf hg hp1) - (λ f g hf hg, hf.add hg) _ hfs - -lemma mem_ℒp.add {f g : α → E} (hf : mem_ℒp f p μ) (hg : mem_ℒp g p μ) : mem_ℒp (f + g) p μ := -⟨ae_strongly_measurable.add hf.1 hg.1, snorm_add_lt_top hf hg⟩ - -lemma mem_ℒp.sub {f g : α → E} (hf : mem_ℒp f p μ) (hg : mem_ℒp g p μ) : mem_ℒp (f - g) p μ := -by { rw sub_eq_add_neg, exact hf.add hg.neg } - -lemma mem_ℒp_finset_sum {ι} (s : finset ι) {f : ι → α → E} (hf : ∀ i ∈ s, mem_ℒp (f i) p μ) : - mem_ℒp (λ a, ∑ i in s, f i a) p μ := -begin - haveI : decidable_eq ι := classical.dec_eq _, - revert hf, - refine finset.induction_on s _ _, - { simp only [zero_mem_ℒp', finset.sum_empty, implies_true_iff], }, - { intros i s his ih hf, - simp only [his, finset.sum_insert, not_false_iff], - 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 - -section monotonicity - -lemma snorm'_le_nnreal_smul_snorm'_of_ae_le_mul {f : α → F} {g : α → G} {c : ℝ≥0} - (h : ∀ᵐ x ∂μ, ‖f x‖₊ ≤ c * ‖g x‖₊) {p : ℝ} (hp : 0 < p) : - snorm' f p μ ≤ c • snorm' g p μ := -begin - simp_rw [snorm'], - rw [←ennreal.rpow_le_rpow_iff hp, ennreal.smul_def, smul_eq_mul, - ennreal.mul_rpow_of_nonneg _ _ hp.le], - simp_rw [←ennreal.rpow_mul, one_div, inv_mul_cancel hp.ne.symm, ennreal.rpow_one, - ennreal.coe_rpow_of_nonneg _ hp.le, ←lintegral_const_mul' _ _ ennreal.coe_ne_top, - ←ennreal.coe_mul], - apply lintegral_mono_ae, - simp_rw [ennreal.coe_le_coe, ←nnreal.mul_rpow, nnreal.rpow_le_rpow_iff hp], - exact h, -end - -lemma snorm_ess_sup_le_nnreal_smul_snorm_ess_sup_of_ae_le_mul {f : α → F} {g : α → G} {c : ℝ≥0} - (h : ∀ᵐ x ∂μ, ‖f x‖₊ ≤ c * ‖g x‖₊) : - snorm_ess_sup f μ ≤ c • snorm_ess_sup g μ := -calc ess_sup (λ x, (‖f x‖₊: ℝ≥0∞)) μ - ≤ ess_sup (λ x, (↑(c * ‖g x‖₊) : ℝ≥0∞)) μ - : ess_sup_mono_ae $ h.mono $ λ x hx, ennreal.coe_le_coe.mpr hx -... = ess_sup (λ x, (c * ‖g x‖₊ : ℝ≥0∞)) μ : by simp_rw ennreal.coe_mul -... = c • ess_sup (λ x, (‖g x‖₊ : ℝ≥0∞)) μ : ennreal.ess_sup_const_mul - -lemma snorm_le_nnreal_smul_snorm_of_ae_le_mul {f : α → F} {g : α → G} {c : ℝ≥0} - (h : ∀ᵐ x ∂μ, ‖f x‖₊ ≤ c * ‖g x‖₊) (p : ℝ≥0∞) : - snorm f p μ ≤ c • snorm g p μ := -begin - by_cases h0 : p = 0, - { simp [h0], }, - by_cases h_top : p = ∞, - { rw h_top, - exact snorm_ess_sup_le_nnreal_smul_snorm_ess_sup_of_ae_le_mul h, }, - simp_rw snorm_eq_snorm' h0 h_top, - exact snorm'_le_nnreal_smul_snorm'_of_ae_le_mul h (ennreal.to_real_pos h0 h_top), -end - --- TODO: add the whole family of lemmas? -private lemma le_mul_iff_eq_zero_of_nonneg_of_neg_of_nonneg {α} [linear_ordered_semiring α] - {a b c : α} (ha : 0 ≤ a) (hb : b < 0) (hc : 0 ≤ c) : a ≤ b * c ↔ a = 0 ∧ c = 0 := -begin - split, - { intro h, - exact ⟨(h.trans (mul_nonpos_of_nonpos_of_nonneg hb.le hc)).antisymm ha, - (nonpos_of_mul_nonneg_right (ha.trans h) hb).antisymm hc⟩ }, - { rintro ⟨rfl, rfl⟩, - rw mul_zero, } -end - -/-- When `c` is negative, `‖f x‖ ≤ c * ‖g x‖` is nonsense and forces both `f` and `g` to have an -`snorm` of `0`. -/ -lemma snorm_eq_zero_and_zero_of_ae_le_mul_neg {f : α → F} {g : α → G} {c : ℝ} - (h : ∀ᵐ x ∂μ, ‖f x‖ ≤ c * ‖g x‖) (hc : c < 0) (p : ℝ≥0∞) : - snorm f p μ = 0 ∧ snorm g p μ = 0 := -begin - simp_rw [le_mul_iff_eq_zero_of_nonneg_of_neg_of_nonneg (norm_nonneg _) hc (norm_nonneg _), - norm_eq_zero, eventually_and] at h, - change f =ᵐ[μ] 0 ∧ g =ᵐ[μ] 0 at h, - simp [snorm_congr_ae h.1, snorm_congr_ae h.2], -end - -lemma snorm_le_mul_snorm_of_ae_le_mul {f : α → F} {g : α → G} {c : ℝ} - (h : ∀ᵐ x ∂μ, ‖f x‖ ≤ c * ‖g x‖) (p : ℝ≥0∞) : - snorm f p μ ≤ (ennreal.of_real c) * snorm g p μ := -snorm_le_nnreal_smul_snorm_of_ae_le_mul - (h.mono $ λ x hx, hx.trans $ mul_le_mul_of_nonneg_right c.le_coe_to_nnreal (norm_nonneg _)) _ - -lemma mem_ℒp.of_nnnorm_le_mul {f : α → E} {g : α → F} {c : ℝ≥0} - (hg : mem_ℒp g p μ) (hf : ae_strongly_measurable f μ) (hfg : ∀ᵐ x ∂μ, ‖f x‖₊ ≤ c * ‖g x‖₊) : - mem_ℒp f p μ := -⟨hf, (snorm_le_nnreal_smul_snorm_of_ae_le_mul hfg p).trans_lt $ - ennreal.mul_lt_top ennreal.coe_ne_top hg.snorm_ne_top⟩ - -lemma mem_ℒp.of_le_mul {f : α → E} {g : α → F} {c : ℝ} - (hg : mem_ℒp g p μ) (hf : ae_strongly_measurable f μ) (hfg : ∀ᵐ x ∂μ, ‖f x‖ ≤ c * ‖g x‖) : - mem_ℒp f p μ := -⟨hf, (snorm_le_mul_snorm_of_ae_le_mul hfg p).trans_lt $ - ennreal.mul_lt_top ennreal.of_real_ne_top hg.snorm_ne_top⟩ - -lemma snorm'_le_snorm'_mul_snorm' {p q r : ℝ} - {f : α → E} (hf : ae_strongly_measurable f μ) {g : α → F} (hg : ae_strongly_measurable g μ) - (b : E → F → G) (h : ∀ᵐ x ∂μ, ‖b (f x) (g x)‖₊ ≤ ‖f x‖₊ * ‖g x‖₊) - (hp0_lt : 0 < p) (hpq : p < q) (hpqr : 1/p = 1/q + 1/r) : - snorm' (λ x, b (f x) (g x)) p μ ≤ snorm' f q μ * snorm' g r μ := -begin - rw snorm', - calc (∫⁻ (a : α), ↑‖b (f a) (g a)‖₊ ^ p ∂μ) ^ (1 / p) - ≤ (∫⁻ (a : α), ↑(‖f a‖₊ * ‖g a‖₊) ^ p ∂μ) ^ (1 / p) : - (ennreal.rpow_le_rpow_iff $ one_div_pos.mpr (hp0_lt)).mpr $ - lintegral_mono_ae $ h.mono $ λ a ha, (ennreal.rpow_le_rpow_iff (hp0_lt)).mpr $ - ennreal.coe_le_coe.mpr $ ha - ... ≤ _ : _, - simp_rw [snorm', ennreal.coe_mul], - exact ennreal.lintegral_Lp_mul_le_Lq_mul_Lr hp0_lt hpq hpqr μ hf.ennnorm - hg.ennnorm, -end - -lemma snorm_le_snorm_top_mul_snorm (p : ℝ≥0∞) - (f : α → E) {g : α → F} (hg : ae_strongly_measurable g μ) (b : E → F → G) - (h : ∀ᵐ x ∂μ, ‖b (f x) (g x)‖₊ ≤ ‖f x‖₊ * ‖g x‖₊) : - snorm (λ x, b (f x) (g x)) p μ ≤ snorm f ∞ μ * snorm g p μ := -begin - by_cases hp_top : p = ∞, - { simp_rw [hp_top, snorm_exponent_top], - refine le_trans (ess_sup_mono_ae $ h.mono $ λ a ha, _) (ennreal.ess_sup_mul_le _ _), - simp_rw [pi.mul_apply, ←ennreal.coe_mul, ennreal.coe_le_coe], - exact ha }, - by_cases hp_zero : p = 0, - { simp only [hp_zero, snorm_exponent_zero, mul_zero, le_zero_iff], }, - simp_rw [snorm_eq_lintegral_rpow_nnnorm hp_zero hp_top, snorm_exponent_top, snorm_ess_sup], - calc (∫⁻ x, ↑‖b (f x) (g x)‖₊ ^ p.to_real ∂μ) ^ (1 / p.to_real) - ≤ (∫⁻ x, ↑‖f x‖₊ ^ p.to_real * ↑‖g x‖₊ ^ p.to_real ∂μ) ^ (1 / p.to_real) : - begin - refine ennreal.rpow_le_rpow _ (one_div_nonneg.mpr ennreal.to_real_nonneg), - refine lintegral_mono_ae (h.mono $ λ a ha, _), - rw ←ennreal.mul_rpow_of_nonneg _ _ ennreal.to_real_nonneg, - refine ennreal.rpow_le_rpow _ ennreal.to_real_nonneg, - rw [←ennreal.coe_mul, ennreal.coe_le_coe], - exact ha, - end - ... ≤ (∫⁻ x, (ess_sup (λ x, ↑‖f x‖₊) μ) ^ p.to_real * ↑‖g x‖₊ ^ p.to_real ∂μ) ^ (1 / p.to_real) : - begin - refine ennreal.rpow_le_rpow _ _, - swap, { rw one_div_nonneg, exact ennreal.to_real_nonneg, }, - refine lintegral_mono_ae _, - filter_upwards [@ennreal.ae_le_ess_sup _ _ μ (λ x, ↑‖f x‖₊)] with x hx, - exact mul_le_mul_right' (ennreal.rpow_le_rpow hx ennreal.to_real_nonneg) _ - end - ... = ess_sup (λ x, ↑‖f x‖₊) μ * (∫⁻ x, ↑‖g x‖₊ ^ p.to_real ∂μ) ^ (1 / p.to_real) : - begin - rw lintegral_const_mul'', - swap, { exact hg.nnnorm.ae_measurable.coe_nnreal_ennreal.pow ae_measurable_const, }, - rw ennreal.mul_rpow_of_nonneg, - swap, { rw one_div_nonneg, exact ennreal.to_real_nonneg, }, - rw [← ennreal.rpow_mul, one_div, mul_inv_cancel, ennreal.rpow_one], - rw [ne.def, ennreal.to_real_eq_zero_iff, auto.not_or_eq], - exact ⟨hp_zero, hp_top⟩, - end -end - -lemma snorm_le_snorm_mul_snorm_top (p : ℝ≥0∞) - {f : α → E} (hf : ae_strongly_measurable f μ) (g : α → F) (b : E → F → G) - (h : ∀ᵐ x ∂μ, ‖b (f x) (g x)‖₊ ≤ ‖f x‖₊ * ‖g x‖₊) : - snorm (λ x, b (f x) (g x)) p μ ≤ snorm f p μ * snorm g ∞ μ := -begin - rw [← snorm_norm f, ← snorm_norm g], - refine (snorm_mono_ae_real h).trans _, - simp_rw [mul_comm ‖f _‖₊, nnreal.coe_mul, coe_nnnorm], - rw mul_comm, - refine snorm_le_snorm_top_mul_snorm p (λ x, ‖g x‖) hf.norm _ (h.mono $ λ x hx, _), - simp_rw [nnnorm_mul], -end - -/-- Hölder's inequality, as an inequality on the `ℒp` seminorm of an elementwise operation -`λ x, b (f x) (g x)`. -/ -lemma snorm_le_snorm_mul_snorm_of_nnnorm {p q r : ℝ≥0∞} - {f : α → E} (hf : ae_strongly_measurable f μ) {g : α → F} (hg : ae_strongly_measurable g μ) - (b : E → F → G) (h : ∀ᵐ x ∂μ, ‖b (f x) (g x)‖₊ ≤ ‖f x‖₊ * ‖g x‖₊) - (hpqr : 1/p = 1/q + 1/r) : - snorm (λ x, b (f x) (g x)) p μ ≤ snorm f q μ * snorm g r μ := -begin - by_cases hp_zero : p = 0, - { simp [hp_zero], }, - have hq_ne_zero : q ≠ 0, - { intro hq_zero, - simp only [hq_zero, hp_zero, one_div, ennreal.inv_zero, top_add, - ennreal.inv_eq_top] at hpqr, - exact hpqr, }, - have hr_ne_zero : r ≠ 0, - { intro hr_zero, - simp only [hr_zero, hp_zero, one_div, ennreal.inv_zero, add_top, - ennreal.inv_eq_top] at hpqr, - exact hpqr, }, - by_cases hq_top : q = ∞, - { have hpr : p = r, - { simpa only [hq_top, one_div, ennreal.div_top, zero_add, inv_inj] using hpqr, }, - rw [← hpr, hq_top], - exact snorm_le_snorm_top_mul_snorm p f hg b h, }, - by_cases hr_top : r = ∞, - { have hpq : p = q, - { simpa only [hr_top, one_div, ennreal.div_top, add_zero, inv_inj] using hpqr, }, - rw [← hpq, hr_top], - exact snorm_le_snorm_mul_snorm_top p hf g b h, }, - have hpq : p < q, - { suffices : 1 / q < 1 / p, - { rwa [one_div, one_div, ennreal.inv_lt_inv] at this, }, - rw hpqr, - refine ennreal.lt_add_right _ _, - { simp only [hq_ne_zero, one_div, ne.def, ennreal.inv_eq_top, not_false_iff], }, - { simp only [hr_top, one_div, ne.def, ennreal.inv_eq_zero, not_false_iff], }, }, - rw [snorm_eq_snorm' hp_zero (hpq.trans_le le_top).ne, snorm_eq_snorm' hq_ne_zero hq_top, - snorm_eq_snorm' hr_ne_zero hr_top], - refine snorm'_le_snorm'_mul_snorm' hf hg _ h _ _ _, - { exact ennreal.to_real_pos hp_zero (hpq.trans_le le_top).ne, }, - { exact ennreal.to_real_strict_mono hq_top hpq, }, - rw [← ennreal.one_to_real, ← ennreal.to_real_div, ← ennreal.to_real_div, ← ennreal.to_real_div, - hpqr, ennreal.to_real_add], - { simp only [hq_ne_zero, one_div, ne.def, ennreal.inv_eq_top, not_false_iff], }, - { simp only [hr_ne_zero, one_div, ne.def, ennreal.inv_eq_top, not_false_iff], }, -end - -/-- Hölder's inequality, as an inequality on the `ℒp` seminorm of an elementwise operation -`λ x, b (f x) (g x)`. -/ -lemma snorm_le_snorm_mul_snorm'_of_norm {p q r : ℝ≥0∞} - {f : α → E} (hf : ae_strongly_measurable f μ) {g : α → F} (hg : ae_strongly_measurable g μ) - (b : E → F → G) (h : ∀ᵐ x ∂μ, ‖b (f x) (g x)‖ ≤ ‖f x‖ * ‖g x‖) - (hpqr : 1/p = 1/q + 1/r) : - snorm (λ x, b (f x) (g x)) p μ ≤ snorm f q μ * snorm g r μ := -snorm_le_snorm_mul_snorm_of_nnnorm hf hg b h hpqr - -end monotonicity - -/-! -### Bounded actions by normed rings - -In this section we show inequalities on the norm. --/ -section has_bounded_smul - -variables {𝕜 : Type*} [normed_ring 𝕜] [mul_action_with_zero 𝕜 E] [mul_action_with_zero 𝕜 F] -variables [has_bounded_smul 𝕜 E] [has_bounded_smul 𝕜 F] - -lemma snorm'_const_smul_le (c : 𝕜) (f : α → F) (hq_pos : 0 < q) : - snorm' (c • f) q μ ≤ ‖c‖₊ • snorm' f q μ := -snorm'_le_nnreal_smul_snorm'_of_ae_le_mul (eventually_of_forall $ λ a, nnnorm_smul_le _ _) hq_pos - -lemma snorm_ess_sup_const_smul_le (c : 𝕜) (f : α → F) : - snorm_ess_sup (c • f) μ ≤ ‖c‖₊ • snorm_ess_sup f μ := -snorm_ess_sup_le_nnreal_smul_snorm_ess_sup_of_ae_le_mul - (eventually_of_forall $ λ a, nnnorm_smul_le _ _) - -lemma snorm_const_smul_le (c : 𝕜) (f : α → F) : - snorm (c • f) p μ ≤ ‖c‖₊ • snorm f p μ := -snorm_le_nnreal_smul_snorm_of_ae_le_mul (eventually_of_forall $ λ a, nnnorm_smul_le _ _) _ - -lemma mem_ℒp.const_smul {f : α → E} (hf : mem_ℒp f p μ) (c : 𝕜) : - mem_ℒp (c • f) p μ := -⟨ae_strongly_measurable.const_smul hf.1 c, - (snorm_const_smul_le c f).trans_lt (ennreal.mul_lt_top ennreal.coe_ne_top hf.2.ne)⟩ - -lemma mem_ℒp.const_mul {R} [normed_ring R] {f : α → R} (hf : mem_ℒp f p μ) (c : R) : - mem_ℒp (λ x, c * f x) p μ := -hf.const_smul c - -lemma snorm'_smul_le_mul_snorm' {p q r : ℝ} - {f : α → E} (hf : ae_strongly_measurable f μ) {φ : α → 𝕜} (hφ : ae_strongly_measurable φ μ) - (hp0_lt : 0 < p) (hpq : p < q) (hpqr : 1/p = 1/q + 1/r) : - snorm' (φ • f) p μ ≤ snorm' φ q μ * snorm' f r μ := -snorm'_le_snorm'_mul_snorm' hφ hf (•) - (eventually_of_forall $ λ a, nnnorm_smul_le _ _) hp0_lt hpq hpqr - -lemma snorm_smul_le_snorm_top_mul_snorm (p : ℝ≥0∞) - {f : α → E} (hf : ae_strongly_measurable f μ) (φ : α → 𝕜) : - snorm (φ • f) p μ ≤ snorm φ ∞ μ * snorm f p μ := -(snorm_le_snorm_top_mul_snorm p φ hf (•) (eventually_of_forall $ λ a, nnnorm_smul_le _ _) : _) - -lemma snorm_smul_le_snorm_mul_snorm_top (p : ℝ≥0∞) - (f : α → E) {φ : α → 𝕜} (hφ : ae_strongly_measurable φ μ) : - snorm (φ • f) p μ ≤ snorm φ p μ * snorm f ∞ μ := -(snorm_le_snorm_mul_snorm_top p hφ f (•) (eventually_of_forall $ λ a, nnnorm_smul_le _ _) : _) - -/-- Hölder's inequality, as an inequality on the `ℒp` seminorm of a scalar product `φ • f`. -/ -lemma snorm_smul_le_mul_snorm {p q r : ℝ≥0∞} - {f : α → E} (hf : ae_strongly_measurable f μ) {φ : α → 𝕜} (hφ : ae_strongly_measurable φ μ) - (hpqr : 1/p = 1/q + 1/r) : - snorm (φ • f) p μ ≤ snorm φ q μ * snorm f r μ := -(snorm_le_snorm_mul_snorm_of_nnnorm hφ hf (•) - (eventually_of_forall $ λ a, nnnorm_smul_le _ _) hpqr : _) - -lemma mem_ℒp.smul {p q r : ℝ≥0∞} {f : α → E} {φ : α → 𝕜} - (hf : mem_ℒp f r μ) (hφ : mem_ℒp φ q μ) (hpqr : 1/p = 1/q + 1/r) : - mem_ℒp (φ • f) p μ := -⟨hφ.1.smul hf.1, (snorm_smul_le_mul_snorm hf.1 hφ.1 hpqr).trans_lt - (ennreal.mul_lt_top hφ.snorm_ne_top hf.snorm_ne_top)⟩ - -lemma mem_ℒp.smul_of_top_right {p : ℝ≥0∞} {f : α → E} {φ : α → 𝕜} - (hf : mem_ℒp f p μ) (hφ : mem_ℒp φ ∞ μ) : - mem_ℒp (φ • f) p μ := -by { apply hf.smul hφ, simp only [ennreal.div_top, zero_add] } - -lemma mem_ℒp.smul_of_top_left {p : ℝ≥0∞} {f : α → E} {φ : α → 𝕜} - (hf : mem_ℒp f ∞ μ) (hφ : mem_ℒp φ p μ) : - mem_ℒp (φ • f) p μ := -by { apply hf.smul hφ, simp only [ennreal.div_top, add_zero] } - -end has_bounded_smul - -/-! -### Bounded actions by normed division rings - -The inequalities in the previous section are now tight. --/ -section normed_space - -variables {𝕜 : Type*} [normed_division_ring 𝕜] [mul_action_with_zero 𝕜 E] [module 𝕜 F] -variables [has_bounded_smul 𝕜 E] [has_bounded_smul 𝕜 F] - -lemma snorm'_const_smul {f : α → F} (c : 𝕜) (hq_pos : 0 < q) : - snorm' (c • f) q μ = ‖c‖₊ • snorm' f q μ := -begin - obtain rfl | hc := eq_or_ne c 0, - { simp [snorm', hq_pos], }, - refine le_antisymm (snorm'_const_smul_le _ _ hq_pos) _, - have : snorm' _ q μ ≤ _:= snorm'_const_smul_le (c⁻¹) (c • f) hq_pos, - rwa [inv_smul_smul₀ hc, nnnorm_inv, ennreal.le_inv_smul_iff (nnnorm_ne_zero_iff.mpr hc)] at this, -end - -lemma snorm_ess_sup_const_smul (c : 𝕜) (f : α → F) : - snorm_ess_sup (c • f) μ = (‖c‖₊ : ℝ≥0∞) * snorm_ess_sup f μ := -by simp_rw [snorm_ess_sup, pi.smul_apply, nnnorm_smul, ennreal.coe_mul, ennreal.ess_sup_const_mul] - -lemma snorm_const_smul (c : 𝕜) (f : α → F) : - snorm (c • f) p μ = (‖c‖₊ : ℝ≥0∞) * snorm f p μ := -begin - obtain rfl | hc := eq_or_ne c 0, - { simp, }, - refine le_antisymm (snorm_const_smul_le _ _) _, - have : snorm _ p μ ≤ _:= snorm_const_smul_le (c⁻¹) (c • f), - rwa [inv_smul_smul₀ hc, nnnorm_inv, ennreal.le_inv_smul_iff (nnnorm_ne_zero_iff.mpr hc)] at this, -end - -end normed_space - -lemma snorm_indicator_ge_of_bdd_below (hp : p ≠ 0) (hp' : p ≠ ∞) - {f : α → F} (C : ℝ≥0) {s : set α} (hs : measurable_set s) - (hf : ∀ᵐ x ∂μ, x ∈ s → C ≤ ‖s.indicator f x‖₊) : - C • μ s ^ (1 / p.to_real) ≤ snorm (s.indicator f) p μ := -begin - rw [ennreal.smul_def, smul_eq_mul, snorm_eq_lintegral_rpow_nnnorm hp hp', - ennreal.le_rpow_one_div_iff (ennreal.to_real_pos hp hp'), - ennreal.mul_rpow_of_nonneg _ _ ennreal.to_real_nonneg, - ← ennreal.rpow_mul, one_div_mul_cancel (ennreal.to_real_pos hp hp').ne.symm, ennreal.rpow_one, - ← set_lintegral_const, ← lintegral_indicator _ hs], - refine lintegral_mono_ae _, - filter_upwards [hf] with x hx, - rw nnnorm_indicator_eq_indicator_nnnorm, - by_cases hxs : x ∈ s, - { simp only [set.indicator_of_mem hxs] at ⊢ hx, - exact ennreal.rpow_le_rpow (ennreal.coe_le_coe.2 (hx hxs)) ennreal.to_real_nonneg }, - { simp [set.indicator_of_not_mem hxs] }, -end - -section is_R_or_C -variables {𝕜 : Type*} [is_R_or_C 𝕜] {f : α → 𝕜} - -lemma mem_ℒp.re (hf : mem_ℒp f p μ) : mem_ℒp (λ x, is_R_or_C.re (f x)) p μ := -begin - have : ∀ x, ‖is_R_or_C.re (f x)‖ ≤ 1 * ‖f x‖, - by { intro x, rw one_mul, exact is_R_or_C.norm_re_le_norm (f x), }, - refine hf.of_le_mul _ (eventually_of_forall this), - exact is_R_or_C.continuous_re.comp_ae_strongly_measurable hf.1, -end - -lemma mem_ℒp.im (hf : mem_ℒp f p μ) : mem_ℒp (λ x, is_R_or_C.im (f x)) p μ := -begin - have : ∀ x, ‖is_R_or_C.im (f x)‖ ≤ 1 * ‖f x‖, - by { intro x, rw one_mul, exact is_R_or_C.norm_im_le_norm (f x), }, - refine hf.of_le_mul _ (eventually_of_forall this), - exact is_R_or_C.continuous_im.comp_ae_strongly_measurable hf.1, -end - -end is_R_or_C - -section liminf - -variables [measurable_space E] [opens_measurable_space E] {R : ℝ≥0} - -lemma ae_bdd_liminf_at_top_rpow_of_snorm_bdd {p : ℝ≥0∞} - {f : ℕ → α → E} (hfmeas : ∀ n, measurable (f n)) (hbdd : ∀ n, snorm (f n) p μ ≤ R) : - ∀ᵐ x ∂μ, liminf (λ n, (‖f n x‖₊ ^ p.to_real : ℝ≥0∞)) at_top < ∞ := -begin - by_cases hp0 : p.to_real = 0, - { simp only [hp0, ennreal.rpow_zero], - refine eventually_of_forall (λ x, _), - rw liminf_const (1 : ℝ≥0∞), - exacts [ennreal.one_lt_top, at_top_ne_bot] }, - have hp : p ≠ 0 := λ h, by simpa [h] using hp0, - have hp' : p ≠ ∞ := λ h, by simpa [h] using hp0, - refine ae_lt_top - (measurable_liminf (λ n, (hfmeas n).nnnorm.coe_nnreal_ennreal.pow_const p.to_real)) - (lt_of_le_of_lt (lintegral_liminf_le - (λ n, (hfmeas n).nnnorm.coe_nnreal_ennreal.pow_const p.to_real)) - (lt_of_le_of_lt _ (ennreal.rpow_lt_top_of_nonneg - ennreal.to_real_nonneg ennreal.coe_ne_top : ↑R ^ p.to_real < ∞))).ne, - simp_rw snorm_eq_lintegral_rpow_nnnorm hp hp' at hbdd, - simp_rw [liminf_eq, eventually_at_top], - exact Sup_le (λ b ⟨a, ha⟩, (ha a le_rfl).trans - ((ennreal.rpow_one_div_le_iff (ennreal.to_real_pos hp hp')).1 (hbdd _))), -end - -lemma ae_bdd_liminf_at_top_of_snorm_bdd {p : ℝ≥0∞} (hp : p ≠ 0) - {f : ℕ → α → E} (hfmeas : ∀ n, measurable (f n)) (hbdd : ∀ n, snorm (f n) p μ ≤ R) : - ∀ᵐ x ∂μ, liminf (λ n, (‖f n x‖₊ : ℝ≥0∞)) at_top < ∞ := -begin - by_cases hp' : p = ∞, - { subst hp', - simp_rw snorm_exponent_top at hbdd, - have : ∀ n, ∀ᵐ x ∂μ, (‖f n x‖₊ : ℝ≥0∞) < R + 1 := - λ n, ae_lt_of_ess_sup_lt (lt_of_le_of_lt (hbdd n) $ - ennreal.lt_add_right ennreal.coe_ne_top one_ne_zero), - rw ← ae_all_iff at this, - filter_upwards [this] with x hx using lt_of_le_of_lt - (liminf_le_of_frequently_le' $ frequently_of_forall $ λ n, (hx n).le) - (ennreal.add_lt_top.2 ⟨ennreal.coe_lt_top, ennreal.one_lt_top⟩) }, - filter_upwards [ae_bdd_liminf_at_top_rpow_of_snorm_bdd hfmeas hbdd] with x hx, - have hppos : 0 < p.to_real := ennreal.to_real_pos hp hp', - have : liminf (λ n, (‖f n x‖₊ ^ p.to_real : ℝ≥0∞)) at_top = - (liminf (λ n, (‖f n x‖₊ : ℝ≥0∞)) at_top)^ p.to_real, - { change liminf (λ n, ennreal.order_iso_rpow p.to_real hppos (‖f n x‖₊ : ℝ≥0∞)) at_top = - ennreal.order_iso_rpow p.to_real hppos (liminf (λ n, (‖f n x‖₊ : ℝ≥0∞)) at_top), - refine (order_iso.liminf_apply (ennreal.order_iso_rpow p.to_real _) _ _ _ _).symm; - is_bounded_default }, - rw this at hx, - rw [← ennreal.rpow_one (liminf (λ n, ‖f n x‖₊) at_top), ← mul_inv_cancel hppos.ne.symm, - ennreal.rpow_mul], - exact ennreal.rpow_lt_top_of_nonneg (inv_nonneg.2 hppos.le) hx.ne, -end - -end liminf - -end ℒp - /-! ### Lp space From 0b7c740e25651db0ba63648fbae9f9d6f941e31b Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Sat, 27 May 2023 15:17:24 +0000 Subject: [PATCH 1096/1291] chore(*): add mathlib4 synchronization comments (#19099) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.category.Group.abelian` * `algebra.category.Group.colimits` * `algebra.category.Group.images` * `algebra.category.Module.abelian` * `algebra.category.Module.biproducts` * `algebra.category.Module.images` * `algebra.category.Ring.constructions` * `algebra.category.Ring.filtered_colimits` * `algebra.category.Ring.limits` * `algebra.module.graded_module` * `algebra.module.torsion` * `algebraic_geometry.presheafed_space.has_colimits` * `algebraic_geometry.prime_spectrum.maximal` * `algebraic_geometry.prime_spectrum.noetherian` * `algebraic_geometry.sheafed_space` * `analysis.box_integral.partition.filter` * `analysis.calculus.conformal.inner_product` * `analysis.calculus.conformal.normed_space` * `analysis.calculus.fderiv.bilinear` * `analysis.calculus.fderiv.mul` * `analysis.complex.arg` * `analysis.complex.conformal` * `analysis.complex.unit_disc.basic` * `analysis.convex.krein_milman` * `analysis.convex.specific_functions.basic` * `analysis.inner_product_space.basic` * `analysis.inner_product_space.conformal_linear_map` * `analysis.inner_product_space.dual` * `analysis.inner_product_space.orthogonal` * `analysis.inner_product_space.projection` * `analysis.inner_product_space.symmetric` * `analysis.locally_convex.abs_convex` * `analysis.locally_convex.weak_dual` * `analysis.mean_inequalities` * `analysis.mean_inequalities_pow` * `analysis.normed_space.dual` * `analysis.normed_space.hahn_banach.separation` * `analysis.normed_space.pi_Lp` * `analysis.normed_space.star.multiplier` * `analysis.normed_space.weak_dual` * `analysis.p_series` * `analysis.special_functions.pow.asymptotics` * `analysis.special_functions.pow.continuity` * `category_theory.abelian.injective` * `category_theory.abelian.injective_resolution` * `category_theory.abelian.projective` * `category_theory.abelian.right_derived` * `category_theory.preadditive.yoneda.limits` * `data.nat.squarefree` * `data.zmod.quotient` * `field_theory.ratfunc` * `geometry.euclidean.angle.unoriented.affine` * `geometry.euclidean.angle.unoriented.basic` * `geometry.euclidean.angle.unoriented.conformal` * `geometry.euclidean.basic` * `geometry.euclidean.inversion` * `geometry.manifold.conformal_groupoid` * `group_theory.exponent` * `group_theory.p_group` * `group_theory.specific_groups.cyclic` * `group_theory.specific_groups.dihedral` * `group_theory.torsion` * `linear_algebra.free_module.ideal_quotient` * `linear_algebra.matrix.bilinear_form` * `linear_algebra.matrix.sesquilinear_form` * `measure_theory.function.egorov` * `measure_theory.function.special_functions.inner` * `measure_theory.function.strongly_measurable.inner` * `measure_theory.integral.mean_inequalities` * `model_theory.graph` * `model_theory.satisfiability` * `model_theory.types` * `model_theory.ultraproducts` * `number_theory.bernoulli` * `number_theory.padics.hensel` * `number_theory.padics.padic_integers` * `order.category.FinPartOrd` * `probability.probability_mass_function.constructions` * `ring_theory.adjoin.field` * `ring_theory.algebraic_independent` * `ring_theory.integral_domain` * `ring_theory.is_tensor_product` * `ring_theory.polynomial.dickson` * `topology.continuous_function.compact` * `topology.instances.complex` * `topology.metric_space.holder` * `topology.sheaves.functors` * `topology.sheaves.sheaf_condition.equalizer_products` * `topology.sheaves.sheaf_condition.pairwise_intersections` --- src/algebra/category/Group/abelian.lean | 3 +++ src/algebra/category/Group/colimits.lean | 3 +++ src/algebra/category/Group/images.lean | 3 +++ src/algebra/category/Module/abelian.lean | 3 +++ src/algebra/category/Module/biproducts.lean | 3 +++ src/algebra/category/Module/images.lean | 3 +++ src/algebra/category/Ring/constructions.lean | 3 +++ src/algebra/category/Ring/filtered_colimits.lean | 3 +++ src/algebra/category/Ring/limits.lean | 3 +++ src/algebra/module/graded_module.lean | 3 +++ src/algebra/module/torsion.lean | 3 +++ src/algebraic_geometry/presheafed_space/has_colimits.lean | 3 +++ src/algebraic_geometry/prime_spectrum/maximal.lean | 3 +++ src/algebraic_geometry/prime_spectrum/noetherian.lean | 3 +++ src/algebraic_geometry/sheafed_space.lean | 3 +++ src/analysis/box_integral/partition/filter.lean | 3 +++ src/analysis/calculus/conformal/inner_product.lean | 3 +++ src/analysis/calculus/conformal/normed_space.lean | 3 +++ src/analysis/calculus/fderiv/bilinear.lean | 3 +++ src/analysis/calculus/fderiv/mul.lean | 3 +++ src/analysis/complex/arg.lean | 3 +++ src/analysis/complex/conformal.lean | 3 +++ src/analysis/complex/unit_disc/basic.lean | 3 +++ src/analysis/convex/krein_milman.lean | 3 +++ src/analysis/convex/specific_functions/basic.lean | 3 +++ src/analysis/inner_product_space/basic.lean | 3 +++ src/analysis/inner_product_space/conformal_linear_map.lean | 3 +++ src/analysis/inner_product_space/dual.lean | 3 +++ src/analysis/inner_product_space/orthogonal.lean | 3 +++ src/analysis/inner_product_space/projection.lean | 3 +++ src/analysis/inner_product_space/symmetric.lean | 3 +++ src/analysis/locally_convex/abs_convex.lean | 3 +++ src/analysis/locally_convex/weak_dual.lean | 3 +++ src/analysis/mean_inequalities.lean | 3 +++ src/analysis/mean_inequalities_pow.lean | 3 +++ src/analysis/normed_space/dual.lean | 3 +++ src/analysis/normed_space/hahn_banach/separation.lean | 3 +++ src/analysis/normed_space/pi_Lp.lean | 3 +++ src/analysis/normed_space/star/multiplier.lean | 3 +++ src/analysis/normed_space/weak_dual.lean | 3 +++ src/analysis/p_series.lean | 3 +++ src/analysis/special_functions/pow/asymptotics.lean | 3 +++ src/analysis/special_functions/pow/continuity.lean | 3 +++ src/category_theory/abelian/injective.lean | 3 +++ src/category_theory/abelian/injective_resolution.lean | 3 +++ src/category_theory/abelian/projective.lean | 3 +++ src/category_theory/abelian/right_derived.lean | 3 +++ src/category_theory/preadditive/yoneda/limits.lean | 3 +++ src/data/nat/squarefree.lean | 3 +++ src/data/zmod/quotient.lean | 3 +++ src/field_theory/ratfunc.lean | 3 +++ src/geometry/euclidean/angle/unoriented/affine.lean | 3 +++ src/geometry/euclidean/angle/unoriented/basic.lean | 3 +++ src/geometry/euclidean/angle/unoriented/conformal.lean | 3 +++ src/geometry/euclidean/basic.lean | 3 +++ src/geometry/euclidean/inversion.lean | 3 +++ src/geometry/manifold/conformal_groupoid.lean | 3 +++ src/group_theory/exponent.lean | 3 +++ src/group_theory/p_group.lean | 3 +++ src/group_theory/specific_groups/cyclic.lean | 3 +++ src/group_theory/specific_groups/dihedral.lean | 3 +++ src/group_theory/torsion.lean | 3 +++ src/linear_algebra/free_module/ideal_quotient.lean | 3 +++ src/linear_algebra/matrix/bilinear_form.lean | 3 +++ src/linear_algebra/matrix/sesquilinear_form.lean | 3 +++ src/measure_theory/function/egorov.lean | 3 +++ src/measure_theory/function/special_functions/inner.lean | 3 +++ src/measure_theory/function/strongly_measurable/inner.lean | 3 +++ src/measure_theory/integral/mean_inequalities.lean | 3 +++ src/model_theory/graph.lean | 3 +++ src/model_theory/satisfiability.lean | 3 +++ src/model_theory/types.lean | 3 +++ src/model_theory/ultraproducts.lean | 3 +++ src/number_theory/bernoulli.lean | 3 +++ src/number_theory/padics/hensel.lean | 3 +++ src/number_theory/padics/padic_integers.lean | 3 +++ src/order/category/FinPartOrd.lean | 3 +++ src/probability/probability_mass_function/constructions.lean | 3 +++ src/ring_theory/adjoin/field.lean | 3 +++ src/ring_theory/algebraic_independent.lean | 3 +++ src/ring_theory/integral_domain.lean | 3 +++ src/ring_theory/is_tensor_product.lean | 3 +++ src/ring_theory/polynomial/dickson.lean | 3 +++ src/topology/continuous_function/compact.lean | 3 +++ src/topology/instances/complex.lean | 3 +++ src/topology/metric_space/holder.lean | 3 +++ src/topology/sheaves/functors.lean | 3 +++ src/topology/sheaves/sheaf_condition/equalizer_products.lean | 3 +++ .../sheaves/sheaf_condition/pairwise_intersections.lean | 3 +++ 89 files changed, 267 insertions(+) diff --git a/src/algebra/category/Group/abelian.lean b/src/algebra/category/Group/abelian.lean index a11ad5a1a383a..757ef5fdc8848 100644 --- a/src/algebra/category/Group/abelian.lean +++ b/src/algebra/category/Group/abelian.lean @@ -11,6 +11,9 @@ import category_theory.abelian.basic /-! # The category of abelian groups is abelian + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ open category_theory diff --git a/src/algebra/category/Group/colimits.lean b/src/algebra/category/Group/colimits.lean index 362f9ade3e506..7853b3466a1d9 100644 --- a/src/algebra/category/Group/colimits.lean +++ b/src/algebra/category/Group/colimits.lean @@ -11,6 +11,9 @@ import category_theory.concrete_category.elementwise /-! # The category of additive commutative groups has all colimits. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file uses a "pre-automated" approach, just as for `Mon/colimits.lean`. It is a very uniform approach, that conceivably could be synthesised directly by a tactic that analyses the shape of `add_comm_group` and `monoid_hom`. diff --git a/src/algebra/category/Group/images.lean b/src/algebra/category/Group/images.lean index 38ddb704cead0..a7c6d9101f729 100644 --- a/src/algebra/category/Group/images.lean +++ b/src/algebra/category/Group/images.lean @@ -9,6 +9,9 @@ import category_theory.limits.shapes.images /-! # The category of commutative additive groups has images. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Note that we don't need to register any of the constructions here as instances, because we get them from the fact that `AddCommGroup` is an abelian category. -/ diff --git a/src/algebra/category/Module/abelian.lean b/src/algebra/category/Module/abelian.lean index ebb1f0aa591da..f2e94e5b62e50 100644 --- a/src/algebra/category/Module/abelian.lean +++ b/src/algebra/category/Module/abelian.lean @@ -11,6 +11,9 @@ import category_theory.abelian.exact /-! # The category of left R-modules is abelian. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Additionally, two linear maps are exact in the categorical sense iff `range f = ker g`. -/ diff --git a/src/algebra/category/Module/biproducts.lean b/src/algebra/category/Module/biproducts.lean index ac7e918c8c678..c03e3fef42f71 100644 --- a/src/algebra/category/Module/biproducts.lean +++ b/src/algebra/category/Module/biproducts.lean @@ -10,6 +10,9 @@ import algebra.homology.short_exact.abelian /-! # The category of `R`-modules has finite biproducts + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ open category_theory diff --git a/src/algebra/category/Module/images.lean b/src/algebra/category/Module/images.lean index 0853e69c1bcb8..8f1145217cd45 100644 --- a/src/algebra/category/Module/images.lean +++ b/src/algebra/category/Module/images.lean @@ -9,6 +9,9 @@ import category_theory.limits.shapes.images /-! # The category of R-modules has images. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Note that we don't need to register any of the constructions here as instances, because we get them from the fact that `Module R` is an abelian category. -/ diff --git a/src/algebra/category/Ring/constructions.lean b/src/algebra/category/Ring/constructions.lean index 5748e6ac7b196..945151b41057f 100644 --- a/src/algebra/category/Ring/constructions.lean +++ b/src/algebra/category/Ring/constructions.lean @@ -13,6 +13,9 @@ import ring_theory.subring.basic /-! # Constructions of (co)limits in CommRing +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we provide the explicit (co)cones for various (co)limits in `CommRing`, including * tensor product is the pushout * `Z` is the initial object diff --git a/src/algebra/category/Ring/filtered_colimits.lean b/src/algebra/category/Ring/filtered_colimits.lean index 342960aea4128..c91d378323f14 100644 --- a/src/algebra/category/Ring/filtered_colimits.lean +++ b/src/algebra/category/Ring/filtered_colimits.lean @@ -9,6 +9,9 @@ import algebra.category.Group.filtered_colimits /-! # The forgetful functor from (commutative) (semi-) rings preserves filtered colimits. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Forgetful functors from algebraic categories usually don't preserve colimits. However, they tend to preserve _filtered_ colimits. diff --git a/src/algebra/category/Ring/limits.lean b/src/algebra/category/Ring/limits.lean index d16259498a862..c59aaa2422e9d 100644 --- a/src/algebra/category/Ring/limits.lean +++ b/src/algebra/category/Ring/limits.lean @@ -11,6 +11,9 @@ import ring_theory.subring.basic /-! # The category of (commutative) rings has all limits +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Further, these limits are preserved by the forgetful functor --- that is, the underlying types are just the limits in the category of types. -/ diff --git a/src/algebra/module/graded_module.lean b/src/algebra/module/graded_module.lean index 085b9dbb85c0d..969e6d008f875 100644 --- a/src/algebra/module/graded_module.lean +++ b/src/algebra/module/graded_module.lean @@ -12,6 +12,9 @@ import algebra.module.big_operators /-! # Graded Module +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given an `R`-algebra `A` graded by `𝓐`, a graded `A`-module `M` is expressed as `direct_sum.decomposition 𝓜` and `set_like.has_graded_smul 𝓐 𝓜`. Then `⨁ i, 𝓜 i` is an `A`-module and is isomorphic to `M`. diff --git a/src/algebra/module/torsion.lean b/src/algebra/module/torsion.lean index 5b2743ad7dcaa..fae52a8083c44 100644 --- a/src/algebra/module/torsion.lean +++ b/src/algebra/module/torsion.lean @@ -13,6 +13,9 @@ import ring_theory.finiteness /-! # Torsion submodules +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions * `torsion_of R M x` : the torsion ideal of `x`, containing all `a` such that `a • x = 0`. diff --git a/src/algebraic_geometry/presheafed_space/has_colimits.lean b/src/algebraic_geometry/presheafed_space/has_colimits.lean index fe3ad8102f4d2..c5bc2df14affb 100644 --- a/src/algebraic_geometry/presheafed_space/has_colimits.lean +++ b/src/algebraic_geometry/presheafed_space/has_colimits.lean @@ -10,6 +10,9 @@ import topology.sheaves.limits /-! # `PresheafedSpace C` has colimits. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + If `C` has limits, then the category `PresheafedSpace C` has colimits, and the forgetful functor to `Top` preserves these colimits. diff --git a/src/algebraic_geometry/prime_spectrum/maximal.lean b/src/algebraic_geometry/prime_spectrum/maximal.lean index 7af181eb1bfb6..4cbe42e75ea5c 100644 --- a/src/algebraic_geometry/prime_spectrum/maximal.lean +++ b/src/algebraic_geometry/prime_spectrum/maximal.lean @@ -10,6 +10,9 @@ import ring_theory.localization.as_subring /-! # Maximal spectrum of a commutative ring +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The maximal spectrum of a commutative ring is the type of all maximal ideals. It is naturally a subset of the prime spectrum endowed with the subspace topology. diff --git a/src/algebraic_geometry/prime_spectrum/noetherian.lean b/src/algebraic_geometry/prime_spectrum/noetherian.lean index 38c3c4e6393e8..6606e553a5a53 100644 --- a/src/algebraic_geometry/prime_spectrum/noetherian.lean +++ b/src/algebraic_geometry/prime_spectrum/noetherian.lean @@ -6,6 +6,9 @@ Authors: Filippo A. E. Nuccio, Andrew Yang import algebraic_geometry.prime_spectrum.basic import topology.noetherian_space /-! +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves additional properties of the prime spectrum a ring is Noetherian. -/ diff --git a/src/algebraic_geometry/sheafed_space.lean b/src/algebraic_geometry/sheafed_space.lean index fb2ce6f310a2e..93836059c50d0 100644 --- a/src/algebraic_geometry/sheafed_space.lean +++ b/src/algebraic_geometry/sheafed_space.lean @@ -9,6 +9,9 @@ import topology.sheaves.functors /-! # Sheafed spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Introduces the category of topological spaces equipped with a sheaf (taking values in an arbitrary target category `C`.) diff --git a/src/analysis/box_integral/partition/filter.lean b/src/analysis/box_integral/partition/filter.lean index d797111834b4a..ce72861841768 100644 --- a/src/analysis/box_integral/partition/filter.lean +++ b/src/analysis/box_integral/partition/filter.lean @@ -9,6 +9,9 @@ import analysis.box_integral.partition.split /-! # Filters used in box-based integrals +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + First we define a structure `box_integral.integration_params`. This structure will be used as an argument in the definition of `box_integral.integral` in order to use the same definition for a few well-known definitions of integrals based on partitions of a rectangular box into subboxes (Riemann diff --git a/src/analysis/calculus/conformal/inner_product.lean b/src/analysis/calculus/conformal/inner_product.lean index 6daefbcfb7ebf..b81bb673c1eda 100644 --- a/src/analysis/calculus/conformal/inner_product.lean +++ b/src/analysis/calculus/conformal/inner_product.lean @@ -9,6 +9,9 @@ import analysis.inner_product_space.conformal_linear_map /-! # Conformal maps between inner product spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A function between inner product spaces is which has a derivative at `x` is conformal at `x` iff the derivative preserves inner products up to a scalar multiple. -/ diff --git a/src/analysis/calculus/conformal/normed_space.lean b/src/analysis/calculus/conformal/normed_space.lean index 04916db1c4a99..3806e1b137282 100644 --- a/src/analysis/calculus/conformal/normed_space.lean +++ b/src/analysis/calculus/conformal/normed_space.lean @@ -12,6 +12,9 @@ import analysis.calculus.fderiv.restrict_scalars /-! # Conformal Maps +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A continuous linear map between real normed spaces `X` and `Y` is `conformal_at` some point `x` if it is real differentiable at that point and its differential `is_conformal_linear_map`. diff --git a/src/analysis/calculus/fderiv/bilinear.lean b/src/analysis/calculus/fderiv/bilinear.lean index e37f6b4dcadf2..2ce933af25aee 100644 --- a/src/analysis/calculus/fderiv/bilinear.lean +++ b/src/analysis/calculus/fderiv/bilinear.lean @@ -8,6 +8,9 @@ import analysis.calculus.fderiv.prod /-! # The derivative of bounded bilinear maps +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + For detailed documentation of the Fréchet derivative, see the module docstring of `analysis/calculus/fderiv/basic.lean`. diff --git a/src/analysis/calculus/fderiv/mul.lean b/src/analysis/calculus/fderiv/mul.lean index 203ae912cf5b0..f52bc203dcd4d 100644 --- a/src/analysis/calculus/fderiv/mul.lean +++ b/src/analysis/calculus/fderiv/mul.lean @@ -8,6 +8,9 @@ import analysis.calculus.fderiv.bilinear /-! # Multiplicative operations on derivatives +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + For detailed documentation of the Fréchet derivative, see the module docstring of `analysis/calculus/fderiv/basic.lean`. diff --git a/src/analysis/complex/arg.lean b/src/analysis/complex/arg.lean index cdebb9f3deb68..3ca878c0a8de3 100644 --- a/src/analysis/complex/arg.lean +++ b/src/analysis/complex/arg.lean @@ -10,6 +10,9 @@ import analysis.special_functions.complex.arg /-! # Rays in the complex numbers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file links the definition `same_ray ℝ x y` with the equality of arguments of complex numbers, the usual way this is considered. diff --git a/src/analysis/complex/conformal.lean b/src/analysis/complex/conformal.lean index 895a1c2cc86c3..849a4a284010e 100644 --- a/src/analysis/complex/conformal.lean +++ b/src/analysis/complex/conformal.lean @@ -10,6 +10,9 @@ import analysis.normed_space.finite_dimension /-! # Conformal maps between complex vector spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We prove the sufficient and necessary conditions for a real-linear map between complex vector spaces to be conformal. diff --git a/src/analysis/complex/unit_disc/basic.lean b/src/analysis/complex/unit_disc/basic.lean index dfcdf4eb3fed6..a93174aad280f 100644 --- a/src/analysis/complex/unit_disc/basic.lean +++ b/src/analysis/complex/unit_disc/basic.lean @@ -9,6 +9,9 @@ import analysis.normed_space.ball_action /-! # Poincaré disc +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define `complex.unit_disc` to be the unit disc in the complex plane. We also introduce some basic operations on this disc. -/ diff --git a/src/analysis/convex/krein_milman.lean b/src/analysis/convex/krein_milman.lean index 4619d869f6420..8ae0817cf2298 100644 --- a/src/analysis/convex/krein_milman.lean +++ b/src/analysis/convex/krein_milman.lean @@ -9,6 +9,9 @@ import analysis.normed_space.hahn_banach.separation /-! # The Krein-Milman theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves the Krein-Milman lemma and the Krein-Milman theorem. ## The lemma diff --git a/src/analysis/convex/specific_functions/basic.lean b/src/analysis/convex/specific_functions/basic.lean index a4ee577d8966b..e79cfe707bc79 100644 --- a/src/analysis/convex/specific_functions/basic.lean +++ b/src/analysis/convex/specific_functions/basic.lean @@ -10,6 +10,9 @@ import tactic.linear_combination /-! # Collection of convex functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove that the following functions are convex or strictly convex: * `strict_convex_on_exp` : The exponential function is strictly convex. diff --git a/src/analysis/inner_product_space/basic.lean b/src/analysis/inner_product_space/basic.lean index 5535f38f68fcd..173e0f0282104 100644 --- a/src/analysis/inner_product_space/basic.lean +++ b/src/analysis/inner_product_space/basic.lean @@ -13,6 +13,9 @@ import linear_algebra.bilinear_form /-! # Inner product space +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines inner product spaces and proves the basic properties. We do not formally define Hilbert spaces, but they can be obtained using the set of assumptions `[normed_add_comm_group E] [inner_product_space 𝕜 E] [complete_space E]`. diff --git a/src/analysis/inner_product_space/conformal_linear_map.lean b/src/analysis/inner_product_space/conformal_linear_map.lean index e252ae03ccca7..a62d2e97b02e0 100644 --- a/src/analysis/inner_product_space/conformal_linear_map.lean +++ b/src/analysis/inner_product_space/conformal_linear_map.lean @@ -9,6 +9,9 @@ import analysis.inner_product_space.basic /-! # Conformal maps between inner product spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In an inner product space, a map is conformal iff it preserves inner products up to a scalar factor. -/ diff --git a/src/analysis/inner_product_space/dual.lean b/src/analysis/inner_product_space/dual.lean index adbbc2ad54936..2ec353f8c15f2 100644 --- a/src/analysis/inner_product_space/dual.lean +++ b/src/analysis/inner_product_space/dual.lean @@ -10,6 +10,9 @@ import analysis.normed_space.star.basic /-! # The Fréchet-Riesz representation theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We consider an inner product space `E` over `𝕜`, which is either `ℝ` or `ℂ`. We define `to_dual_map`, a conjugate-linear isometric embedding of `E` into its dual, which maps an element `x` of the space to `λ y, ⟪x, y⟫`. diff --git a/src/analysis/inner_product_space/orthogonal.lean b/src/analysis/inner_product_space/orthogonal.lean index e97459a37a164..e82e720f6a709 100644 --- a/src/analysis/inner_product_space/orthogonal.lean +++ b/src/analysis/inner_product_space/orthogonal.lean @@ -9,6 +9,9 @@ import analysis.inner_product_space.basic /-! # Orthogonal complements of submodules +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, the `orthogonal` complement of a submodule `K` is defined, and basic API established. Some of the more subtle results about the orthogonal complement are delayed to `analysis.inner_product_space.projection`. diff --git a/src/analysis/inner_product_space/projection.lean b/src/analysis/inner_product_space/projection.lean index 42e7c5024ac76..3e772563a5854 100644 --- a/src/analysis/inner_product_space/projection.lean +++ b/src/analysis/inner_product_space/projection.lean @@ -13,6 +13,9 @@ import data.is_R_or_C.lemmas /-! # The orthogonal projection +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given a nonempty complete subspace `K` of an inner product space `E`, this file constructs `orthogonal_projection K : E →L[𝕜] K`, the orthogonal projection of `E` onto `K`. This map satisfies: for any point `u` in `E`, the point `v = orthogonal_projection K u` in `K` minimizes the diff --git a/src/analysis/inner_product_space/symmetric.lean b/src/analysis/inner_product_space/symmetric.lean index 4033741a0e5b3..19ad96eb1dc37 100644 --- a/src/analysis/inner_product_space/symmetric.lean +++ b/src/analysis/inner_product_space/symmetric.lean @@ -10,6 +10,9 @@ import linear_algebra.sesquilinear_form /-! # Symmetric linear maps in an inner product space +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines and proves basic theorems about symmetric **not necessarily bounded** operators on an inner product space, i.e linear maps `T : E → E` such that `∀ x y, ⟪T x, y⟫ = ⟪x, T y⟫`. diff --git a/src/analysis/locally_convex/abs_convex.lean b/src/analysis/locally_convex/abs_convex.lean index 2b12ac870128a..b302e21cb2d3c 100644 --- a/src/analysis/locally_convex/abs_convex.lean +++ b/src/analysis/locally_convex/abs_convex.lean @@ -10,6 +10,9 @@ import analysis.convex.gauge /-! # Absolutely convex sets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A set is called absolutely convex or disked if it is convex and balanced. The importance of absolutely convex sets comes from the fact that every locally convex topological vector space has a basis consisting of absolutely convex sets. diff --git a/src/analysis/locally_convex/weak_dual.lean b/src/analysis/locally_convex/weak_dual.lean index 5d601abe1bc30..7a93c7c14106e 100644 --- a/src/analysis/locally_convex/weak_dual.lean +++ b/src/analysis/locally_convex/weak_dual.lean @@ -10,6 +10,9 @@ import analysis.locally_convex.with_seminorms /-! # Weak Dual in Topological Vector Spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We prove that the weak topology induced by a bilinear form `B : E →ₗ[𝕜] F →ₗ[𝕜] 𝕜` is locally convex and we explicit give a neighborhood basis in terms of the family of seminorms `λ x, ‖B x y‖` for `y : F`. diff --git a/src/analysis/mean_inequalities.lean b/src/analysis/mean_inequalities.lean index 4dba0800e08bd..62cbe6a4393f3 100644 --- a/src/analysis/mean_inequalities.lean +++ b/src/analysis/mean_inequalities.lean @@ -11,6 +11,9 @@ import data.real.conjugate_exponents /-! # Mean value inequalities +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove several inequalities for finite sums, including AM-GM inequality, Young's inequality, Hölder inequality, and Minkowski inequality. Versions for integrals of some of these inequalities are available in `measure_theory.mean_inequalities`. diff --git a/src/analysis/mean_inequalities_pow.lean b/src/analysis/mean_inequalities_pow.lean index a694c87faf4f9..f63af0fd6e1de 100644 --- a/src/analysis/mean_inequalities_pow.lean +++ b/src/analysis/mean_inequalities_pow.lean @@ -11,6 +11,9 @@ import tactic.positivity /-! # Mean value inequalities +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove several mean inequalities for finite sums. Versions for integrals of some of these inequalities are available in `measure_theory.mean_inequalities`. diff --git a/src/analysis/normed_space/dual.lean b/src/analysis/normed_space/dual.lean index e2b6c94a46c82..ab0cb02b407cc 100644 --- a/src/analysis/normed_space/dual.lean +++ b/src/analysis/normed_space/dual.lean @@ -10,6 +10,9 @@ import analysis.locally_convex.polar /-! # The topological dual of a normed space +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define the topological dual `normed_space.dual` of a normed space, and the continuous linear map `normed_space.inclusion_in_double_dual` from a normed space into its double dual. diff --git a/src/analysis/normed_space/hahn_banach/separation.lean b/src/analysis/normed_space/hahn_banach/separation.lean index ec6836e48ff2b..42850fe414803 100644 --- a/src/analysis/normed_space/hahn_banach/separation.lean +++ b/src/analysis/normed_space/hahn_banach/separation.lean @@ -11,6 +11,9 @@ import topology.algebra.module.locally_convex /-! # Separation Hahn-Banach theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove the geometric Hahn-Banach theorem. For any two disjoint convex sets, there exists a continuous linear functional separating them, geometrically meaning that we can intercalate a plane between them. diff --git a/src/analysis/normed_space/pi_Lp.lean b/src/analysis/normed_space/pi_Lp.lean index dc735f9d255b2..f74d084aefc29 100644 --- a/src/analysis/normed_space/pi_Lp.lean +++ b/src/analysis/normed_space/pi_Lp.lean @@ -9,6 +9,9 @@ import linear_algebra.matrix.basis /-! # `L^p` distance on finite products of metric spaces + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. Given finitely many metric spaces, one can put the max distance on their product, but there is also a whole family of natural distances, indexed by a parameter `p : ℝ≥0∞`, that also induce the product topology. We define them in this file. For `0 < p < ∞`, the distance on `Π i, α i` diff --git a/src/analysis/normed_space/star/multiplier.lean b/src/analysis/normed_space/star/multiplier.lean index 16a4dd5d64409..5c3efde31d567 100644 --- a/src/analysis/normed_space/star/multiplier.lean +++ b/src/analysis/normed_space/star/multiplier.lean @@ -13,6 +13,9 @@ import analysis.normed_space.star.mul /-! # Multiplier Algebra of a C⋆-algebra +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Define the multiplier algebra of a C⋆-algebra as the algebra (over `𝕜`) of double centralizers, for which we provide the localized notation `𝓜(𝕜, A)`. A double centralizer is a pair of continuous linear maps `L R : A →L[𝕜] A` satisfying the intertwining condition `R x * y = x * L y`. diff --git a/src/analysis/normed_space/weak_dual.lean b/src/analysis/normed_space/weak_dual.lean index a5046c3b02086..2ead0b9970dee 100644 --- a/src/analysis/normed_space/weak_dual.lean +++ b/src/analysis/normed_space/weak_dual.lean @@ -10,6 +10,9 @@ import analysis.normed_space.operator_norm /-! # Weak dual of normed space +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Let `E` be a normed space over a field `𝕜`. This file is concerned with properties of the weak-* topology on the dual of `E`. By the dual, we mean either of the type synonyms `normed_space.dual 𝕜 E` or `weak_dual 𝕜 E`, depending on whether it is viewed as equipped with its diff --git a/src/analysis/p_series.lean b/src/analysis/p_series.lean index ee72f8505e915..20ecaa558a71a 100644 --- a/src/analysis/p_series.lean +++ b/src/analysis/p_series.lean @@ -8,6 +8,9 @@ import analysis.special_functions.pow.nnreal /-! # Convergence of `p`-series +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove that the series `∑' k in ℕ, 1 / k ^ p` converges if and only if `p > 1`. The proof is based on the [Cauchy condensation test](https://en.wikipedia.org/wiki/Cauchy_condensation_test): `∑ k, f k` diff --git a/src/analysis/special_functions/pow/asymptotics.lean b/src/analysis/special_functions/pow/asymptotics.lean index a906c4f2d5db1..e6efb51740b87 100644 --- a/src/analysis/special_functions/pow/asymptotics.lean +++ b/src/analysis/special_functions/pow/asymptotics.lean @@ -9,6 +9,9 @@ import analysis.special_functions.pow.nnreal /-! # Limits and asymptotics of power functions at `+∞` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains results about the limiting behaviour of power functions at `+∞`. For convenience some results on asymptotics as `x → 0` (those which are not just continuity statements) are also located here. diff --git a/src/analysis/special_functions/pow/continuity.lean b/src/analysis/special_functions/pow/continuity.lean index a970291facd82..b59cb77b44b81 100644 --- a/src/analysis/special_functions/pow/continuity.lean +++ b/src/analysis/special_functions/pow/continuity.lean @@ -9,6 +9,9 @@ import analysis.special_functions.pow.asymptotics /-! # Continuity of power functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains lemmas about continuity of the power functions on `ℂ`, `ℝ`, `ℝ≥0`, and `ℝ≥0∞`. -/ diff --git a/src/category_theory/abelian/injective.lean b/src/category_theory/abelian/injective.lean index bf89fc5c7b7ca..781dbd2ab645b 100644 --- a/src/category_theory/abelian/injective.lean +++ b/src/category_theory/abelian/injective.lean @@ -12,6 +12,9 @@ import category_theory.preadditive.yoneda.injective /-! # Injective objects in abelian categories +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + * Objects in an abelian categories are injective if and only if the preadditive Yoneda functor on them preserves finite colimits. -/ diff --git a/src/category_theory/abelian/injective_resolution.lean b/src/category_theory/abelian/injective_resolution.lean index fa70a18e89435..e5fc4068f4dcb 100644 --- a/src/category_theory/abelian/injective_resolution.lean +++ b/src/category_theory/abelian/injective_resolution.lean @@ -10,6 +10,9 @@ import algebra.homology.homotopy_category /-! # Main result +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + When the underlying category is abelian: * `category_theory.InjectiveResolution.desc`: Given `I : InjectiveResolution X` and `J : InjectiveResolution Y`, any morphism `X ⟶ Y` admits a descent to a chain map diff --git a/src/category_theory/abelian/projective.lean b/src/category_theory/abelian/projective.lean index 68156ef0055fd..d24b4877fcff4 100644 --- a/src/category_theory/abelian/projective.lean +++ b/src/category_theory/abelian/projective.lean @@ -11,6 +11,9 @@ import category_theory.preadditive.yoneda.projective /-! # Abelian categories with enough projectives have projective resolutions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + When `C` is abelian `projective.d f` and `f` are exact. Hence, starting from an epimorphism `P ⟶ X`, where `P` is projective, we can apply `projective.d` repeatedly to obtain a projective resolution of `X`. diff --git a/src/category_theory/abelian/right_derived.lean b/src/category_theory/abelian/right_derived.lean index 4797700212e96..e3d6b76f63860 100644 --- a/src/category_theory/abelian/right_derived.lean +++ b/src/category_theory/abelian/right_derived.lean @@ -12,6 +12,9 @@ import category_theory.abelian.exact /-! # Right-derived functors +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define the right-derived functors `F.right_derived n : C ⥤ D` for any additive functor `F` out of a category with injective resolutions. diff --git a/src/category_theory/preadditive/yoneda/limits.lean b/src/category_theory/preadditive/yoneda/limits.lean index 919782c556834..ebbec2599df58 100644 --- a/src/category_theory/preadditive/yoneda/limits.lean +++ b/src/category_theory/preadditive/yoneda/limits.lean @@ -9,6 +9,9 @@ import algebra.category.Module.abelian /-! # The Yoneda embedding for preadditive categories preserves limits +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The Yoneda embedding for preadditive categories preserves limits. ## Implementation notes diff --git a/src/data/nat/squarefree.lean b/src/data/nat/squarefree.lean index 08150e41ae27c..55c412691488f 100644 --- a/src/data/nat/squarefree.lean +++ b/src/data/nat/squarefree.lean @@ -10,6 +10,9 @@ import ring_theory.int.basic /-! # Lemmas about squarefreeness of natural numbers + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. A number is squarefree when it is not divisible by any squares except the squares of units. ## Main Results diff --git a/src/data/zmod/quotient.lean b/src/data/zmod/quotient.lean index 3b4fedf4093b5..c619573eabdc4 100644 --- a/src/data/zmod/quotient.lean +++ b/src/data/zmod/quotient.lean @@ -11,6 +11,9 @@ import ring_theory.ideal.quotient_operations /-! # `zmod n` and quotient groups / rings +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file relates `zmod n` to the quotient group `quotient_add_group.quotient (add_subgroup.zmultiples n)` and to the quotient ring `(ideal.span {n}).quotient`. diff --git a/src/field_theory/ratfunc.lean b/src/field_theory/ratfunc.lean index e796b7f2ec61e..07e44dd53cf0b 100644 --- a/src/field_theory/ratfunc.lean +++ b/src/field_theory/ratfunc.lean @@ -12,6 +12,9 @@ import ring_theory.polynomial.content /-! # The field of rational functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the field `ratfunc K` of rational functions over a field `K`, and shows it is the field of fractions of `K[X]`. diff --git a/src/geometry/euclidean/angle/unoriented/affine.lean b/src/geometry/euclidean/angle/unoriented/affine.lean index 8b4036ae625b1..8a91034cff029 100644 --- a/src/geometry/euclidean/angle/unoriented/affine.lean +++ b/src/geometry/euclidean/angle/unoriented/affine.lean @@ -9,6 +9,9 @@ import geometry.euclidean.angle.unoriented.basic /-! # Angles between points +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines unoriented angles in Euclidean affine spaces. ## Main definitions diff --git a/src/geometry/euclidean/angle/unoriented/basic.lean b/src/geometry/euclidean/angle/unoriented/basic.lean index d332c358a4407..e001f6f43dc7c 100644 --- a/src/geometry/euclidean/angle/unoriented/basic.lean +++ b/src/geometry/euclidean/angle/unoriented/basic.lean @@ -9,6 +9,9 @@ import analysis.special_functions.trigonometric.inverse /-! # Angles between vectors +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines unoriented angles in real inner product spaces. ## Main definitions diff --git a/src/geometry/euclidean/angle/unoriented/conformal.lean b/src/geometry/euclidean/angle/unoriented/conformal.lean index d19eca5cf0dd6..ed2de536eddc9 100644 --- a/src/geometry/euclidean/angle/unoriented/conformal.lean +++ b/src/geometry/euclidean/angle/unoriented/conformal.lean @@ -9,6 +9,9 @@ import geometry.euclidean.angle.unoriented.basic /-! # Angles and conformal maps +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves that conformal maps preserve angles. -/ diff --git a/src/geometry/euclidean/basic.lean b/src/geometry/euclidean/basic.lean index 543b4470efd8d..f8d4cf7ee4f1f 100644 --- a/src/geometry/euclidean/basic.lean +++ b/src/geometry/euclidean/basic.lean @@ -9,6 +9,9 @@ import algebra.quadratic_discriminant /-! # Euclidean spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file makes some definitions and proves very basic geometrical results about real inner product spaces and Euclidean affine spaces. Results about real inner product spaces that involve the norm and diff --git a/src/geometry/euclidean/inversion.lean b/src/geometry/euclidean/inversion.lean index ff0d2dcf46fec..92cd13a99fa17 100644 --- a/src/geometry/euclidean/inversion.lean +++ b/src/geometry/euclidean/inversion.lean @@ -8,6 +8,9 @@ import analysis.inner_product_space.basic /-! # Inversion in an affine space +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define inversion in a sphere in an affine space. This map sends each point `x` to the point `y` such that `y -ᵥ c = (R / dist x c) ^ 2 • (x -ᵥ c)`, where `c` and `R` are the center and the radius the sphere. diff --git a/src/geometry/manifold/conformal_groupoid.lean b/src/geometry/manifold/conformal_groupoid.lean index 7ea9bc0fe5c17..6a75afe7a26f1 100644 --- a/src/geometry/manifold/conformal_groupoid.lean +++ b/src/geometry/manifold/conformal_groupoid.lean @@ -9,6 +9,9 @@ import geometry.manifold.charted_space /-! # Conformal Groupoid +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define the groupoid of conformal maps on normed spaces. ## Main definitions diff --git a/src/group_theory/exponent.lean b/src/group_theory/exponent.lean index f77a742a26313..76bfdd346d6c5 100644 --- a/src/group_theory/exponent.lean +++ b/src/group_theory/exponent.lean @@ -13,6 +13,9 @@ import tactic.by_contra /-! # Exponent of a group +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the exponent of a group, or more generally a monoid. For a group `G` it is defined to be the minimal `n≥1` such that `g ^ n = 1` for all `g ∈ G`. For a finite group `G`, it is equal to the lowest common multiple of the order of all elements of the group `G`. diff --git a/src/group_theory/p_group.lean b/src/group_theory/p_group.lean index c07f252b002d1..0a4c8224a0536 100644 --- a/src/group_theory/p_group.lean +++ b/src/group_theory/p_group.lean @@ -15,6 +15,9 @@ import tactic.interval_cases /-! # p-groups +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains a proof that if `G` is a `p`-group acting on a finite set `α`, then the number of fixed points of the action is congruent mod `p` to the cardinality of `α`. It also contains proofs of some corollaries of this lemma about existence of fixed points. diff --git a/src/group_theory/specific_groups/cyclic.lean b/src/group_theory/specific_groups/cyclic.lean index 9a48297f65a54..c86f41fd18079 100644 --- a/src/group_theory/specific_groups/cyclic.lean +++ b/src/group_theory/specific_groups/cyclic.lean @@ -14,6 +14,9 @@ import group_theory.exponent /-! # Cyclic groups +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A group `G` is called cyclic if there exists an element `g : G` such that every element of `G` is of the form `g ^ n` for some `n : ℕ`. This file only deals with the predicate on a group to be cyclic. For the concrete cyclic group of order `n`, see `data.zmod.basic`. diff --git a/src/group_theory/specific_groups/dihedral.lean b/src/group_theory/specific_groups/dihedral.lean index 7be44e835e223..889f538502344 100644 --- a/src/group_theory/specific_groups/dihedral.lean +++ b/src/group_theory/specific_groups/dihedral.lean @@ -9,6 +9,9 @@ import group_theory.exponent /-! # Dihedral Groups +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define the dihedral groups `dihedral_group n`, with elements `r i` and `sr i` for `i : zmod n`. For `n ≠ 0`, `dihedral_group n` represents the symmetry group of the regular `n`-gon. `r i` diff --git a/src/group_theory/torsion.lean b/src/group_theory/torsion.lean index 964c95393eead..343ef3612fce0 100644 --- a/src/group_theory/torsion.lean +++ b/src/group_theory/torsion.lean @@ -13,6 +13,9 @@ import group_theory.submonoid.operations /-! # Torsion groups +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines torsion groups, i.e. groups where all elements have finite order. ## Main definitions diff --git a/src/linear_algebra/free_module/ideal_quotient.lean b/src/linear_algebra/free_module/ideal_quotient.lean index ab6d4937ceef9..7d1ec9fc5f5ef 100644 --- a/src/linear_algebra/free_module/ideal_quotient.lean +++ b/src/linear_algebra/free_module/ideal_quotient.lean @@ -11,6 +11,9 @@ import linear_algebra.quotient_pi /-! # Ideals in free modules over PIDs +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main results - `ideal.quotient_equiv_pi_span`: `S ⧸ I`, if `S` is finite free as a module over a PID `R`, diff --git a/src/linear_algebra/matrix/bilinear_form.lean b/src/linear_algebra/matrix/bilinear_form.lean index 9e198c8f1b5b0..ea6878ba4f3d3 100644 --- a/src/linear_algebra/matrix/bilinear_form.lean +++ b/src/linear_algebra/matrix/bilinear_form.lean @@ -14,6 +14,9 @@ import linear_algebra.matrix.sesquilinear_form /-! # Bilinear form +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the conversion between bilinear forms and matrices. ## Main definitions diff --git a/src/linear_algebra/matrix/sesquilinear_form.lean b/src/linear_algebra/matrix/sesquilinear_form.lean index 53f690ff9cb04..380bb004a4175 100644 --- a/src/linear_algebra/matrix/sesquilinear_form.lean +++ b/src/linear_algebra/matrix/sesquilinear_form.lean @@ -14,6 +14,9 @@ import linear_algebra.sesquilinear_form /-! # Sesquilinear form +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the conversion between sesquilinear forms and matrices. ## Main definitions diff --git a/src/measure_theory/function/egorov.lean b/src/measure_theory/function/egorov.lean index 63bfb80eb9b2c..d6d2f6c83c226 100644 --- a/src/measure_theory/function/egorov.lean +++ b/src/measure_theory/function/egorov.lean @@ -8,6 +8,9 @@ import measure_theory.function.strongly_measurable.basic /-! # Egorov theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains the Egorov theorem which states that an almost everywhere convergent sequence on a finite measure space converges uniformly except on an arbitrarily small set. This theorem is useful for the Vitali convergence theorem as well as theorems regarding diff --git a/src/measure_theory/function/special_functions/inner.lean b/src/measure_theory/function/special_functions/inner.lean index 1f11270445b53..7ef44069c3dce 100644 --- a/src/measure_theory/function/special_functions/inner.lean +++ b/src/measure_theory/function/special_functions/inner.lean @@ -9,6 +9,9 @@ import measure_theory.constructions.borel_space.complex /-! # Measurability of scalar products + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ variables {α : Type*} {𝕜 : Type*} {E : Type*} diff --git a/src/measure_theory/function/strongly_measurable/inner.lean b/src/measure_theory/function/strongly_measurable/inner.lean index 6b25e4682327c..2d147cf26fd68 100644 --- a/src/measure_theory/function/strongly_measurable/inner.lean +++ b/src/measure_theory/function/strongly_measurable/inner.lean @@ -9,6 +9,9 @@ import analysis.inner_product_space.basic /-! # Inner products of strongly measurable functions are strongly measurable. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + -/ variables {α : Type*} diff --git a/src/measure_theory/integral/mean_inequalities.lean b/src/measure_theory/integral/mean_inequalities.lean index 9c8942b2eba36..970d8d1085a23 100644 --- a/src/measure_theory/integral/mean_inequalities.lean +++ b/src/measure_theory/integral/mean_inequalities.lean @@ -11,6 +11,9 @@ import measure_theory.function.special_functions.basic /-! # Mean value inequalities for integrals +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove several inequalities on integrals, notably the Hölder inequality and the Minkowski inequality. The versions for finite sums are in `analysis.mean_inequalities`. diff --git a/src/model_theory/graph.lean b/src/model_theory/graph.lean index e5d2763808457..f95115deb3c5f 100644 --- a/src/model_theory/graph.lean +++ b/src/model_theory/graph.lean @@ -8,6 +8,9 @@ import combinatorics.simple_graph.basic /-! # First-Ordered Structures in Graph Theory + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. This file defines first-order languages, structures, and theories in graph theory. ## Main Definitions diff --git a/src/model_theory/satisfiability.lean b/src/model_theory/satisfiability.lean index 1c1a693e6e45b..11d17f341a5f3 100644 --- a/src/model_theory/satisfiability.lean +++ b/src/model_theory/satisfiability.lean @@ -9,6 +9,9 @@ import model_theory.skolem /-! # First-Order Satisfiability + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. This file deals with the satisfiability of first-order theories, as well as equivalence over them. ## Main Definitions diff --git a/src/model_theory/types.lean b/src/model_theory/types.lean index 2df5368d2a6f2..bcc49d9b00d11 100644 --- a/src/model_theory/types.lean +++ b/src/model_theory/types.lean @@ -7,6 +7,9 @@ import model_theory.satisfiability /-! # Type Spaces + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. This file defines the space of complete types over a first-order theory. (Note that types in model theory are different from types in type theory.) diff --git a/src/model_theory/ultraproducts.lean b/src/model_theory/ultraproducts.lean index 4c9363e172200..e8608d43a9b42 100644 --- a/src/model_theory/ultraproducts.lean +++ b/src/model_theory/ultraproducts.lean @@ -9,6 +9,9 @@ import order.filter.ultrafilter /-! # Ultraproducts and Łoś's Theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main Definitions * `first_order.language.ultraproduct.Structure` is the ultraproduct structure on `filter.product`. diff --git a/src/number_theory/bernoulli.lean b/src/number_theory/bernoulli.lean index b77fb4c739448..6e204582ab37a 100644 --- a/src/number_theory/bernoulli.lean +++ b/src/number_theory/bernoulli.lean @@ -12,6 +12,9 @@ import tactic.field_simp /-! # Bernoulli numbers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The Bernoulli numbers are a sequence of rational numbers that frequently show up in number theory. diff --git a/src/number_theory/padics/hensel.lean b/src/number_theory/padics/hensel.lean index 57521f1b1a33d..e541fd41c24d7 100644 --- a/src/number_theory/padics/hensel.lean +++ b/src/number_theory/padics/hensel.lean @@ -12,6 +12,9 @@ import topology.metric_space.cau_seq_filter /-! # Hensel's lemma on ℤ_p +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves Hensel's lemma on ℤ_p, roughly following Keith Conrad's writeup: diff --git a/src/number_theory/padics/padic_integers.lean b/src/number_theory/padics/padic_integers.lean index aac1d3115ce20..d0c66151065c3 100644 --- a/src/number_theory/padics/padic_integers.lean +++ b/src/number_theory/padics/padic_integers.lean @@ -9,6 +9,9 @@ import ring_theory.discrete_valuation_ring.basic /-! # p-adic integers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the `p`-adic integers `ℤ_[p]` as the subtype of `ℚ_[p]` with norm `≤ 1`. We show that `ℤ_[p]` * is complete, diff --git a/src/order/category/FinPartOrd.lean b/src/order/category/FinPartOrd.lean index ffa0dab0797f1..228a947fccd7a 100644 --- a/src/order/category/FinPartOrd.lean +++ b/src/order/category/FinPartOrd.lean @@ -9,6 +9,9 @@ import order.category.PartOrd /-! # The category of finite partial orders +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This defines `FinPartOrd`, the category of finite partial orders. Note: `FinPartOrd` is *not* a subcategory of `BddOrd` because finite orders are not necessarily diff --git a/src/probability/probability_mass_function/constructions.lean b/src/probability/probability_mass_function/constructions.lean index ae7df339f0f93..4d3a42989695d 100644 --- a/src/probability/probability_mass_function/constructions.lean +++ b/src/probability/probability_mass_function/constructions.lean @@ -8,6 +8,9 @@ import probability.probability_mass_function.monad /-! # Specific Constructions of Probability Mass Functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file gives a number of different `pmf` constructions for common probability distributions. `map` and `seq` allow pushing a `pmf α` along a function `f : α → β` (or distribution of diff --git a/src/ring_theory/adjoin/field.lean b/src/ring_theory/adjoin/field.lean index 34c9f07d450db..7a4b4920535d2 100644 --- a/src/ring_theory/adjoin/field.lean +++ b/src/ring_theory/adjoin/field.lean @@ -11,6 +11,9 @@ import ring_theory.adjoin_root /-! # Adjoining elements to a field +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Some lemmas on the ring generating by adjoining an element to a field. ## Main statements diff --git a/src/ring_theory/algebraic_independent.lean b/src/ring_theory/algebraic_independent.lean index e7b6cc5ce00c3..e6dbfbfd6aaaa 100644 --- a/src/ring_theory/algebraic_independent.lean +++ b/src/ring_theory/algebraic_independent.lean @@ -12,6 +12,9 @@ import data.mv_polynomial.equiv /-! # Algebraic Independence +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines algebraic independence of a family of element of an `R` algebra ## Main definitions diff --git a/src/ring_theory/integral_domain.lean b/src/ring_theory/integral_domain.lean index ff73d6a5b1cc9..1f325173541f9 100644 --- a/src/ring_theory/integral_domain.lean +++ b/src/ring_theory/integral_domain.lean @@ -11,6 +11,9 @@ import algebra.geom_sum /-! # Integral domains +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Assorted theorems about integral domains. ## Main theorems diff --git a/src/ring_theory/is_tensor_product.lean b/src/ring_theory/is_tensor_product.lean index dceb63548b86c..23e7bdcfeb9b8 100644 --- a/src/ring_theory/is_tensor_product.lean +++ b/src/ring_theory/is_tensor_product.lean @@ -10,6 +10,9 @@ import algebra.module.ulift /-! # The characteristice predicate of tensor product +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions - `is_tensor_product`: A predicate on `f : M₁ →ₗ[R] M₂ →ₗ[R] M` expressing that `f` realizes `M` as diff --git a/src/ring_theory/polynomial/dickson.lean b/src/ring_theory/polynomial/dickson.lean index 0fa68573f549b..c6cb3c157b108 100644 --- a/src/ring_theory/polynomial/dickson.lean +++ b/src/ring_theory/polynomial/dickson.lean @@ -12,6 +12,9 @@ import ring_theory.ideal.local_ring /-! # Dickson polynomials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The (generalised) Dickson polynomials are a family of polynomials indexed by `ℕ × ℕ`, with coefficients in a commutative ring `R` depending on an element `a∈R`. More precisely, the they satisfy the recursion `dickson k a (n + 2) = X * (dickson k a n + 1) - a * (dickson k a n)` diff --git a/src/topology/continuous_function/compact.lean b/src/topology/continuous_function/compact.lean index 7b83f26c73a94..fbf02f9ec680b 100644 --- a/src/topology/continuous_function/compact.lean +++ b/src/topology/continuous_function/compact.lean @@ -11,6 +11,9 @@ import topology.sets.compacts /-! # Continuous functions on a compact space +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Continuous functions `C(α, β)` from a compact space `α` to a metric space `β` are automatically bounded, and so acquire various structures inherited from `α →ᵇ β`. diff --git a/src/topology/instances/complex.lean b/src/topology/instances/complex.lean index bf183b83e7a88..91ac9ecd0bd55 100644 --- a/src/topology/instances/complex.lean +++ b/src/topology/instances/complex.lean @@ -10,6 +10,9 @@ import topology.algebra.uniform_ring /-! # Some results about the topology of ℂ + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ section complex_subfield diff --git a/src/topology/metric_space/holder.lean b/src/topology/metric_space/holder.lean index d7dbd7c8f7cb7..ab2bd573e5f5c 100644 --- a/src/topology/metric_space/holder.lean +++ b/src/topology/metric_space/holder.lean @@ -9,6 +9,9 @@ import analysis.special_functions.pow.continuity /-! # Hölder continuous functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define Hölder continuity on a set and on the whole space. We also prove some basic properties of Hölder continuous functions. diff --git a/src/topology/sheaves/functors.lean b/src/topology/sheaves/functors.lean index 8693dff4d83a3..35679f3981c7e 100644 --- a/src/topology/sheaves/functors.lean +++ b/src/topology/sheaves/functors.lean @@ -9,6 +9,9 @@ import topology.sheaves.sheaf_condition.pairwise_intersections /-! # functors between categories of sheaves +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Show that the pushforward of a sheaf is a sheaf, and define the pushforward functor from the category of C-valued sheaves on X to that of sheaves on Y, given a continuous map between diff --git a/src/topology/sheaves/sheaf_condition/equalizer_products.lean b/src/topology/sheaves/sheaf_condition/equalizer_products.lean index eb5b673ba8e0d..e75084455b939 100644 --- a/src/topology/sheaves/sheaf_condition/equalizer_products.lean +++ b/src/topology/sheaves/sheaf_condition/equalizer_products.lean @@ -10,6 +10,9 @@ import topology.sheaves.sheaf_condition.pairwise_intersections /-! # The sheaf condition in terms of an equalizer of products +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Here we set up the machinery for the "usual" definition of the sheaf condition, e.g. as in https://stacks.math.columbia.edu/tag/0072 in terms of an equalizer diagram where the two objects are diff --git a/src/topology/sheaves/sheaf_condition/pairwise_intersections.lean b/src/topology/sheaves/sheaf_condition/pairwise_intersections.lean index 7ad353b49cde6..93c0cf1430110 100644 --- a/src/topology/sheaves/sheaf_condition/pairwise_intersections.lean +++ b/src/topology/sheaves/sheaf_condition/pairwise_intersections.lean @@ -14,6 +14,9 @@ import algebra.category.Ring.constructions /-! # Equivalent formulations of the sheaf condition +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We give an equivalent formulation of the sheaf condition. Given any indexed type `ι`, we define `overlap ι`, From 5a2df4cd59cb31e97a516d4603a14bed5c2f9425 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sat, 27 May 2023 17:03:32 +0000 Subject: [PATCH 1097/1291] chore(measure_theory/function/simple_func_dense_lp): generalize typeclasses (#19111) Lemmas like `to_Lp_smul` now apply for `normed_ring`s, not just `normed_field`s. --- .../function/simple_func_dense_lp.lean | 22 +++++++++++++------ 1 file changed, 15 insertions(+), 7 deletions(-) diff --git a/src/measure_theory/function/simple_func_dense_lp.lean b/src/measure_theory/function/simple_func_dense_lp.lean index 1853843cf77f6..0c9b2cf2b6b9a 100644 --- a/src/measure_theory/function/simple_func_dense_lp.lean +++ b/src/measure_theory/function/simple_func_dense_lp.lean @@ -445,7 +445,7 @@ unnecessary. But instead, `Lp.simple_func E p μ` is defined as an `add_subgrou which does not permit this (but has the advantage of working when `E` itself is a normed group, i.e. has no scalar action). -/ -variables [normed_field 𝕜] [normed_space 𝕜 E] +variables [normed_ring 𝕜] [module 𝕜 E] [has_bounded_smul 𝕜 E] /-- If `E` is a normed space, `Lp.simple_func E p μ` is a `has_smul`. Not declared as an instance as it is (as of writing) used only in the construction of the Bochner integral. -/ @@ -477,12 +477,20 @@ local attribute [instance] simple_func.module /-- If `E` is a normed space, `Lp.simple_func E p μ` is a normed space. Not declared as an instance as it is (as of writing) used only in the construction of the Bochner integral. -/ -protected def normed_space [fact (1 ≤ p)] : normed_space 𝕜 (Lp.simple_func E p μ) := -⟨ λc f, by { rw [add_subgroup.coe_norm, add_subgroup.coe_norm, coe_smul, norm_smul] } ⟩ +protected lemma has_bounded_smul [fact (1 ≤ p)] : has_bounded_smul 𝕜 (Lp.simple_func E p μ) := +has_bounded_smul.of_norm_smul_le $ λ r f, (norm_smul_le r (f : Lp E p μ) : _) + +local attribute [instance] simple_func.has_bounded_smul + +/-- If `E` is a normed space, `Lp.simple_func E p μ` is a normed space. Not declared as an +instance as it is (as of writing) used only in the construction of the Bochner integral. -/ +protected def normed_space {𝕜} [normed_field 𝕜] [normed_space 𝕜 E] [fact (1 ≤ p)] : + normed_space 𝕜 (Lp.simple_func E p μ) := +⟨norm_smul_le⟩ end instances -local attribute [instance] simple_func.module simple_func.normed_space +local attribute [instance] simple_func.module simple_func.normed_space simple_func.has_bounded_smul section to_Lp @@ -508,7 +516,7 @@ lemma to_Lp_sub (f g : α →ₛ E) (hf : mem_ℒp f p μ) (hg : mem_ℒp g p μ to_Lp (f - g) (hf.sub hg) = to_Lp f hf - to_Lp g hg := by { simp only [sub_eq_add_neg, ← to_Lp_neg, ← to_Lp_add], refl } -variables [normed_field 𝕜] [normed_space 𝕜 E] +variables [normed_ring 𝕜] [module 𝕜 E] [has_bounded_smul 𝕜 E] lemma to_Lp_smul (f : α →ₛ E) (hf : mem_ℒp f p μ) (c : 𝕜) : to_Lp (c • f) (hf.const_smul c) = c • to_Lp f hf := rfl @@ -601,7 +609,7 @@ begin repeat { assume h, rw h, }, end -variables [normed_field 𝕜] [normed_space 𝕜 E] +variables [normed_ring 𝕜] [module 𝕜 E] [has_bounded_smul 𝕜 E] lemma smul_to_simple_func (k : 𝕜) (f : Lp.simple_func E p μ) : to_simple_func (k • f) =ᵐ[μ] k • to_simple_func f := @@ -712,7 +720,7 @@ protected lemma dense_range (hp_ne_top : p ≠ ∞) : dense_range (coe : (Lp.simple_func E p μ) → (Lp E p μ)) := (simple_func.dense_inducing hp_ne_top).dense -variables [normed_field 𝕜] [normed_space 𝕜 E] +variables [normed_ring 𝕜] [module 𝕜 E] [has_bounded_smul 𝕜 E] variables (α E 𝕜) /-- The embedding of Lp simple functions into Lp functions, as a continuous linear map. -/ From 13b0d72fd8533ba459ac66e9a885e35ffabb32b2 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sat, 27 May 2023 18:17:13 +0000 Subject: [PATCH 1098/1291] feat(measure_theory/function/l1_space): generalize multiplicative results to is_unit (#19073) This generalization lets us use these lemmas for left- and right- multiplication by non-commutative normed rings. This doesn't make it all the way to `integral_mul_const`, but it's a step in that direction. --- src/measure_theory/function/l1_space.lean | 62 +++++++++++-------- .../integral/peak_function.lean | 2 +- 2 files changed, 36 insertions(+), 28 deletions(-) diff --git a/src/measure_theory/function/l1_space.lean b/src/measure_theory/function/l1_space.lean index 5a403b13399f0..c0716c5777ff1 100644 --- a/src/measure_theory/function/l1_space.lean +++ b/src/measure_theory/function/l1_space.lean @@ -352,8 +352,8 @@ end pos_part section normed_space variables {𝕜 : Type*} -lemma has_finite_integral.smul' [has_smul 𝕜 β] [has_nnnorm 𝕜] (c : 𝕜) {f : α → β} - (h : ∀ (k : 𝕜) (b : β), ‖k • b‖₊ ≤ ‖k‖₊ * ‖b‖₊) : +lemma has_finite_integral.smul + [normed_add_comm_group 𝕜] [smul_zero_class 𝕜 β] [has_bounded_smul 𝕜 β] (c : 𝕜) {f : α → β} : has_finite_integral f μ → has_finite_integral (c • f) μ := begin simp only [has_finite_integral], assume hfi, @@ -361,7 +361,7 @@ begin ∫⁻ (a : α), ‖c • f a‖₊ ∂μ ≤ ∫⁻ (a : α), (‖c‖₊) * ‖f a‖₊ ∂μ : begin refine lintegral_mono _, intro i, - exact_mod_cast h c (f i), + exact_mod_cast (nnnorm_smul_le c (f i) : _), end ... < ∞ : begin @@ -370,30 +370,28 @@ begin end end -lemma has_finite_integral.smul [normed_field 𝕜] [normed_space 𝕜 β] (c : 𝕜) {f : α → β} : - has_finite_integral f μ → has_finite_integral (c • f) μ := -has_finite_integral.smul' _ $ λ a b, nnnorm_smul_le a b - -lemma has_finite_integral_smul_iff [normed_field 𝕜] [normed_space 𝕜 β] {c : 𝕜} (hc : c ≠ 0) +lemma has_finite_integral_smul_iff + [normed_ring 𝕜] [mul_action_with_zero 𝕜 β] [has_bounded_smul 𝕜 β] + {c : 𝕜} (hc : is_unit c) (f : α → β) : has_finite_integral (c • f) μ ↔ has_finite_integral f μ := begin + obtain ⟨c, rfl⟩ := hc, split, { assume h, - simpa only [smul_smul, inv_mul_cancel hc, one_smul] using h.smul c⁻¹ }, + simpa only [smul_smul, units.inv_mul, one_smul] using h.smul (↑c⁻¹ : 𝕜) }, exact has_finite_integral.smul _ end lemma has_finite_integral.const_mul [normed_ring 𝕜] {f : α → 𝕜} (h : has_finite_integral f μ) (c : 𝕜) : has_finite_integral (λ x, c * f x) μ := -(has_finite_integral.smul' c nnnorm_mul_le h : _) +h.smul c lemma has_finite_integral.mul_const [normed_ring 𝕜] {f : α → 𝕜} (h : has_finite_integral f μ) (c : 𝕜) : has_finite_integral (λ x, f x * c) μ := -(has_finite_integral.smul' (mul_opposite.op c) - (λ a b, (nnnorm_mul_le b a.unop).trans_eq $ mul_comm _ _) h : _) +h.smul (mul_opposite.op c) end normed_space @@ -673,7 +671,8 @@ end /-- Hölder's inequality for integrable functions: the scalar multiplication of an integrable scalar-valued function by a vector-value function with finite essential supremum is integrable. -/ -lemma integrable.smul_ess_sup {𝕜 : Type*} [normed_field 𝕜] [normed_space 𝕜 β] {f : α → 𝕜} +lemma integrable.smul_ess_sup {𝕜 : Type*} [normed_ring 𝕜] [module 𝕜 β] [has_bounded_smul 𝕜 β] + {f : α → 𝕜} (hf : integrable f μ) {g : α → β} (g_ae_strongly_measurable : ae_strongly_measurable g μ) (ess_sup_g : ess_sup (λ x, (‖g x‖₊ : ℝ≥0∞)) μ ≠ ∞) : integrable (λ (x : α), f x • g x) μ := @@ -932,16 +931,25 @@ hf.neg.pos_part end pos_part -section normed_space -variables {𝕜 : Type*} [normed_field 𝕜] [normed_space 𝕜 β] +section has_bounded_smul +variables {𝕜 : Type*} -lemma integrable.smul (c : 𝕜) {f : α → β} +lemma integrable.smul [normed_add_comm_group 𝕜] [smul_zero_class 𝕜 β] [has_bounded_smul 𝕜 β] + (c : 𝕜) {f : α → β} (hf : integrable f μ) : integrable (c • f) μ := ⟨hf.ae_strongly_measurable.const_smul c, hf.has_finite_integral.smul c⟩ -lemma integrable_smul_iff {c : 𝕜} (hc : c ≠ 0) (f : α → β) : +lemma is_unit.integrable_smul_iff [normed_ring 𝕜] [module 𝕜 β] [has_bounded_smul 𝕜 β] + {c : 𝕜} (hc : is_unit c) (f : α → β) : + integrable (c • f) μ ↔ integrable f μ := +and_congr (hc.ae_strongly_measurable_const_smul_iff) (has_finite_integral_smul_iff hc f) + +lemma integrable_smul_iff [normed_division_ring 𝕜] [module 𝕜 β] [has_bounded_smul 𝕜 β] + {c : 𝕜} (hc : c ≠ 0) (f : α → β) : integrable (c • f) μ ↔ integrable f μ := -and_congr (ae_strongly_measurable_const_smul_iff₀ hc) (has_finite_integral_smul_iff hc f) +(is_unit.mk0 _ hc).integrable_smul_iff f + +variables [normed_ring 𝕜] [module 𝕜 β] [has_bounded_smul 𝕜 β] lemma integrable.smul_of_top_right {f : α → β} {φ : α → 𝕜} (hf : integrable f μ) (hφ : mem_ℒp φ ∞ μ) : @@ -957,7 +965,7 @@ lemma integrable.smul_const {f : α → 𝕜} (hf : integrable f μ) (c : β) : integrable (λ x, f x • c) μ := hf.smul_of_top_left (mem_ℒp_top_const c) -end normed_space +end has_bounded_smul section normed_space_over_complete_field variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] [complete_space 𝕜] @@ -980,7 +988,7 @@ variables {𝕜 : Type*} [normed_ring 𝕜] {f : α → 𝕜} lemma integrable.const_mul {f : α → 𝕜} (h : integrable f μ) (c : 𝕜) : integrable (λ x, c * f x) μ := -⟨h.ae_strongly_measurable.const_smul c, h.has_finite_integral.const_mul c⟩ +h.smul c lemma integrable.const_mul' {f : α → 𝕜} (h : integrable f μ) (c : 𝕜) : integrable ((λ (x : α), c) * f) μ := @@ -988,7 +996,7 @@ integrable.const_mul h c lemma integrable.mul_const {f : α → 𝕜} (h : integrable f μ) (c : 𝕜) : integrable (λ x, f x * c) μ := -⟨h.ae_strongly_measurable.const_smul (mul_opposite.op c), h.has_finite_integral.mul_const c⟩ +h.smul (mul_opposite.op c) lemma integrable.mul_const' {f : α → 𝕜} (h : integrable f μ) (c : 𝕜) : integrable (f * (λ (x : α), c)) μ := @@ -996,11 +1004,11 @@ integrable.mul_const h c lemma integrable_const_mul_iff {c : 𝕜} (hc : is_unit c) (f : α → 𝕜) : integrable (λ x, c * f x) μ ↔ integrable f μ := -let ⟨u, hc⟩ := hc in hc ▸ ⟨λ h, by simpa using h.const_mul ↑(u⁻¹), λ h, h.const_mul _⟩ +hc.integrable_smul_iff f lemma integrable_mul_const_iff {c : 𝕜} (hc : is_unit c) (f : α → 𝕜) : integrable (λ x, f x * c) μ ↔ integrable f μ := -let ⟨u, hc⟩ := hc in hc ▸ ⟨λ h, by simpa using h.mul_const ↑(u⁻¹), λ h, h.mul_const _⟩ +hc.op.integrable_smul_iff f lemma integrable.bdd_mul' {f g : α → 𝕜} {c : ℝ} (hg : integrable g μ) (hf : ae_strongly_measurable f μ) (hf_bound : ∀ᵐ x ∂μ, ‖f x‖ ≤ c) : @@ -1136,13 +1144,13 @@ lemma integrable.sub {f g : α →ₘ[μ] β} (hf : integrable f) (hg : integrab end -section normed_space -variables {𝕜 : Type*} [normed_field 𝕜] [normed_space 𝕜 β] +section has_bounded_smul +variables {𝕜 : Type*} [normed_ring 𝕜] [module 𝕜 β] [has_bounded_smul 𝕜 β] lemma integrable.smul {c : 𝕜} {f : α →ₘ[μ] β} : integrable f → integrable (c • f) := induction_on f $ λ f hfm hfi, (integrable_mk _).2 $ ((integrable_mk hfm).1 hfi).smul _ -end normed_space +end has_bounded_smul end @@ -1263,7 +1271,7 @@ by { simp [integrable.to_L1, snorm, snorm'], simp [edist_eq_coe_nnnorm_sub] } edist (hf.to_L1 f) 0 = ∫⁻ a, edist (f a) 0 ∂μ := by { simp [integrable.to_L1, snorm, snorm'], simp [edist_eq_coe_nnnorm] } -variables {𝕜 : Type*} [normed_field 𝕜] [normed_space 𝕜 β] +variables {𝕜 : Type*} [normed_ring 𝕜] [module 𝕜 β] [has_bounded_smul 𝕜 β] lemma to_L1_smul (f : α → β) (hf : integrable f μ) (k : 𝕜) : to_L1 (λ a, k • f a) (hf.smul k) = k • to_L1 f hf := rfl diff --git a/src/measure_theory/integral/peak_function.lean b/src/measure_theory/integral/peak_function.lean index ac8e5df9b3b59..27c8bb7266666 100644 --- a/src/measure_theory/integral/peak_function.lean +++ b/src/measure_theory/integral/peak_function.lean @@ -58,7 +58,7 @@ begin filter_upwards [tendsto_uniformly_on_iff.1 (hlφ u u_open x₀u) 1 zero_lt_one, hiφ] with i hi h'i, have A : integrable_on (λ x, φ i x • g x) (s \ u) μ, - { apply integrable.smul_of_top_right (hmg.mono (diff_subset _ _) le_rfl), + { refine integrable.smul_of_top_right (hmg.mono (diff_subset _ _) le_rfl) _, apply mem_ℒp_top_of_bound ((integrable_of_integral_eq_one h'i).ae_strongly_measurable.mono_set ((diff_subset _ _))) 1, filter_upwards [self_mem_ae_restrict (hs.diff u_open.measurable_set)] with x hx, From 3bce8d800a6f2b8f63fe1e588fd76a9ff4adcebe Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sat, 27 May 2023 19:40:00 +0000 Subject: [PATCH 1099/1291] chore(analysis/calculus/deriv): split file (#19113) * Split `analysis.calculus.deriv` into smaller files. * Add copyright headers and (stubs of) module docstrings. * Change proofs of derivatives of `pow` and `polynomial` so that `analysis.calculus.deriv.pow` does not depend on polynomials. Co-authored-by: Parcly Taxel Co-authored-by: Eric Wieser --- src/analysis/bounded_variation.lean | 4 +- .../calculus/bump_function_inner.lean | 1 + src/analysis/calculus/cont_diff.lean | 1 + src/analysis/calculus/deriv.lean | 2314 ----------------- src/analysis/calculus/deriv/add.lean | 304 +++ src/analysis/calculus/deriv/basic.lean | 599 +++++ src/analysis/calculus/deriv/comp.lean | 263 ++ src/analysis/calculus/deriv/inv.lean | 218 ++ src/analysis/calculus/deriv/inverse.lean | 120 + src/analysis/calculus/deriv/linear.lean | 83 + src/analysis/calculus/deriv/mul.lean | 398 +++ src/analysis/calculus/deriv/polynomial.lean | 157 ++ src/analysis/calculus/deriv/pow.lean | 91 + src/analysis/calculus/deriv/prod.lean | 98 + src/analysis/calculus/deriv/slope.lean | 178 ++ src/analysis/calculus/deriv/star.lean | 66 + src/analysis/calculus/deriv/support.lean | 44 + src/analysis/calculus/deriv/zpow.lean | 149 ++ src/analysis/calculus/diff_cont_on_cl.lean | 2 +- src/analysis/calculus/dslope.lean | 3 +- src/analysis/calculus/fderiv_analytic.lean | 2 +- src/analysis/calculus/fderiv_measurable.lean | 2 +- src/analysis/calculus/iterated_deriv.lean | 2 +- src/analysis/calculus/lhopital.lean | 1 + src/analysis/calculus/local_extr.lean | 2 +- src/analysis/calculus/mean_value.lean | 1 + src/analysis/calculus/monotone.lean | 2 +- src/analysis/complex/real_deriv.lean | 1 + .../convex/specific_functions/deriv.lean | 1 + src/analysis/special_functions/log/deriv.lean | 2 + src/analysis/special_functions/pow/deriv.lean | 1 + src/geometry/manifold/instances/sphere.lean | 1 + .../integral/circle_integral.lean | 1 + .../integral/divergence_theorem.lean | 2 +- .../integral/fund_thm_calculus.lean | 4 + .../polynomial/hermite/gaussian.lean | 3 +- test/simp_command.lean | 1 + 37 files changed, 2798 insertions(+), 2324 deletions(-) delete mode 100644 src/analysis/calculus/deriv.lean create mode 100644 src/analysis/calculus/deriv/add.lean create mode 100644 src/analysis/calculus/deriv/basic.lean create mode 100644 src/analysis/calculus/deriv/comp.lean create mode 100644 src/analysis/calculus/deriv/inv.lean create mode 100644 src/analysis/calculus/deriv/inverse.lean create mode 100644 src/analysis/calculus/deriv/linear.lean create mode 100644 src/analysis/calculus/deriv/mul.lean create mode 100644 src/analysis/calculus/deriv/polynomial.lean create mode 100644 src/analysis/calculus/deriv/pow.lean create mode 100644 src/analysis/calculus/deriv/prod.lean create mode 100644 src/analysis/calculus/deriv/slope.lean create mode 100644 src/analysis/calculus/deriv/star.lean create mode 100644 src/analysis/calculus/deriv/support.lean create mode 100644 src/analysis/calculus/deriv/zpow.lean diff --git a/src/analysis/bounded_variation.lean b/src/analysis/bounded_variation.lean index 98b2b765cfdf7..1fae789f4cb82 100644 --- a/src/analysis/bounded_variation.lean +++ b/src/analysis/bounded_variation.lean @@ -3,10 +3,12 @@ Copyright (c) 2022 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ +import analysis.calculus.deriv.add +import analysis.calculus.fderiv.equiv +import analysis.calculus.fderiv.prod import analysis.calculus.monotone import data.set.function import algebra.group.basic -import tactic.swap_var import tactic.wlog /-! diff --git a/src/analysis/calculus/bump_function_inner.lean b/src/analysis/calculus/bump_function_inner.lean index 2afda5effcfb0..4dcda6d14946a 100644 --- a/src/analysis/calculus/bump_function_inner.lean +++ b/src/analysis/calculus/bump_function_inner.lean @@ -3,6 +3,7 @@ Copyright (c) 2020 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel, Floris van Doorn -/ +import analysis.calculus.deriv.inv import analysis.calculus.extend_deriv import analysis.calculus.iterated_deriv import analysis.inner_product_space.calculus diff --git a/src/analysis/calculus/cont_diff.lean b/src/analysis/calculus/cont_diff.lean index dd1f41175b230..21556cd0fda2b 100644 --- a/src/analysis/calculus/cont_diff.lean +++ b/src/analysis/calculus/cont_diff.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel, Floris van Doorn -/ import analysis.calculus.cont_diff_def +import analysis.calculus.deriv.inverse import analysis.calculus.mean_value import analysis.normed_space.finite_dimension import data.nat.choose.cast diff --git a/src/analysis/calculus/deriv.lean b/src/analysis/calculus/deriv.lean deleted file mode 100644 index dbcd04015b9bb..0000000000000 --- a/src/analysis/calculus/deriv.lean +++ /dev/null @@ -1,2314 +0,0 @@ -/- -Copyright (c) 2019 Gabriel Ebner. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Gabriel Ebner, Sébastien Gouëzel --/ -import analysis.calculus.fderiv.add -import analysis.calculus.fderiv.mul -import analysis.calculus.fderiv.equiv -import analysis.calculus.fderiv.restrict_scalars -import analysis.calculus.fderiv.star -import data.polynomial.algebra_map -import data.polynomial.derivative -import linear_algebra.affine_space.slope - -/-! - -# One-dimensional derivatives - -This file defines the derivative of a function `f : 𝕜 → F` where `𝕜` is a -normed field and `F` is a normed space over this field. The derivative of -such a function `f` at a point `x` is given by an element `f' : F`. - -The theory is developed analogously to the [Fréchet -derivatives](./fderiv.html). We first introduce predicates defined in terms -of the corresponding predicates for Fréchet derivatives: - - - `has_deriv_at_filter f f' x L` states that the function `f` has the - derivative `f'` at the point `x` as `x` goes along the filter `L`. - - - `has_deriv_within_at f f' s x` states that the function `f` has the - derivative `f'` at the point `x` within the subset `s`. - - - `has_deriv_at f f' x` states that the function `f` has the derivative `f'` - at the point `x`. - - - `has_strict_deriv_at f f' x` states that the function `f` has the derivative `f'` - at the point `x` in the sense of strict differentiability, i.e., - `f y - f z = (y - z) • f' + o (y - z)` as `y, z → x`. - -For the last two notions we also define a functional version: - - - `deriv_within f s x` is a derivative of `f` at `x` within `s`. If the - derivative does not exist, then `deriv_within f s x` equals zero. - - - `deriv f x` is a derivative of `f` at `x`. If the derivative does not - exist, then `deriv f x` equals zero. - -The theorems `fderiv_within_deriv_within` and `fderiv_deriv` show that the -one-dimensional derivatives coincide with the general Fréchet derivatives. - -We also show the existence and compute the derivatives of: - - constants - - the identity function - - linear maps - - addition - - sum of finitely many functions - - negation - - subtraction - - multiplication - - star - - inverse `x → x⁻¹` - - multiplication of two functions in `𝕜 → 𝕜` - - multiplication of a function in `𝕜 → 𝕜` and of a function in `𝕜 → E` - - composition of a function in `𝕜 → F` with a function in `𝕜 → 𝕜` - - composition of a function in `F → E` with a function in `𝕜 → F` - - inverse function (assuming that it exists; the inverse function theorem is in `inverse.lean`) - - division - - polynomials - -For most binary operations we also define `const_op` and `op_const` theorems for the cases when -the first or second argument is a constant. This makes writing chains of `has_deriv_at`'s easier, -and they more frequently lead to the desired result. - -We set up the simplifier so that it can compute the derivative of simple functions. For instance, -```lean -example (x : ℝ) : deriv (λ x, cos (sin x) * exp x) x = (cos(sin(x))-sin(sin(x))*cos(x))*exp(x) := -by { simp, ring } -``` - -## Implementation notes - -Most of the theorems are direct restatements of the corresponding theorems -for Fréchet derivatives. - -The strategy to construct simp lemmas that give the simplifier the possibility to compute -derivatives is the same as the one for differentiability statements, as explained in `fderiv.lean`. -See the explanations there. --/ - -universes u v w -noncomputable theory -open_locale classical topology big_operators filter ennreal polynomial -open filter asymptotics set -open continuous_linear_map (smul_right smul_right_one_eq_iff) - - -variables {𝕜 : Type u} [nontrivially_normed_field 𝕜] - -section -variables {F : Type v} [normed_add_comm_group F] [normed_space 𝕜 F] -variables {E : Type w} [normed_add_comm_group E] [normed_space 𝕜 E] - -/-- -`f` has the derivative `f'` at the point `x` as `x` goes along the filter `L`. - -That is, `f x' = f x + (x' - x) • f' + o(x' - x)` where `x'` converges along the filter `L`. --/ -def has_deriv_at_filter (f : 𝕜 → F) (f' : F) (x : 𝕜) (L : filter 𝕜) := -has_fderiv_at_filter f (smul_right (1 : 𝕜 →L[𝕜] 𝕜) f') x L - -/-- -`f` has the derivative `f'` at the point `x` within the subset `s`. - -That is, `f x' = f x + (x' - x) • f' + o(x' - x)` where `x'` converges to `x` inside `s`. --/ -def has_deriv_within_at (f : 𝕜 → F) (f' : F) (s : set 𝕜) (x : 𝕜) := -has_deriv_at_filter f f' x (𝓝[s] x) - -/-- -`f` has the derivative `f'` at the point `x`. - -That is, `f x' = f x + (x' - x) • f' + o(x' - x)` where `x'` converges to `x`. --/ -def has_deriv_at (f : 𝕜 → F) (f' : F) (x : 𝕜) := -has_deriv_at_filter f f' x (𝓝 x) - -/-- `f` has the derivative `f'` at the point `x` in the sense of strict differentiability. - -That is, `f y - f z = (y - z) • f' + o(y - z)` as `y, z → x`. -/ -def has_strict_deriv_at (f : 𝕜 → F) (f' : F) (x : 𝕜) := -has_strict_fderiv_at f (smul_right (1 : 𝕜 →L[𝕜] 𝕜) f') x - -/-- -Derivative of `f` at the point `x` within the set `s`, if it exists. Zero otherwise. - -If the derivative exists (i.e., `∃ f', has_deriv_within_at f f' s x`), then -`f x' = f x + (x' - x) • deriv_within f s x + o(x' - x)` where `x'` converges to `x` inside `s`. --/ -def deriv_within (f : 𝕜 → F) (s : set 𝕜) (x : 𝕜) := -fderiv_within 𝕜 f s x 1 - -/-- -Derivative of `f` at the point `x`, if it exists. Zero otherwise. - -If the derivative exists (i.e., `∃ f', has_deriv_at f f' x`), then -`f x' = f x + (x' - x) • deriv f x + o(x' - x)` where `x'` converges to `x`. --/ -def deriv (f : 𝕜 → F) (x : 𝕜) := -fderiv 𝕜 f x 1 - -variables {f f₀ f₁ g : 𝕜 → F} -variables {f' f₀' f₁' g' : F} -variables {x : 𝕜} -variables {s t : set 𝕜} -variables {L L₁ L₂ : filter 𝕜} - -/-- Expressing `has_fderiv_at_filter f f' x L` in terms of `has_deriv_at_filter` -/ -lemma has_fderiv_at_filter_iff_has_deriv_at_filter {f' : 𝕜 →L[𝕜] F} : - has_fderiv_at_filter f f' x L ↔ has_deriv_at_filter f (f' 1) x L := -by simp [has_deriv_at_filter] - -lemma has_fderiv_at_filter.has_deriv_at_filter {f' : 𝕜 →L[𝕜] F} : - has_fderiv_at_filter f f' x L → has_deriv_at_filter f (f' 1) x L := -has_fderiv_at_filter_iff_has_deriv_at_filter.mp - -/-- Expressing `has_fderiv_within_at f f' s x` in terms of `has_deriv_within_at` -/ -lemma has_fderiv_within_at_iff_has_deriv_within_at {f' : 𝕜 →L[𝕜] F} : - has_fderiv_within_at f f' s x ↔ has_deriv_within_at f (f' 1) s x := -has_fderiv_at_filter_iff_has_deriv_at_filter - -/-- Expressing `has_deriv_within_at f f' s x` in terms of `has_fderiv_within_at` -/ -lemma has_deriv_within_at_iff_has_fderiv_within_at {f' : F} : - has_deriv_within_at f f' s x ↔ - has_fderiv_within_at f (smul_right (1 : 𝕜 →L[𝕜] 𝕜) f') s x := -iff.rfl - -lemma has_fderiv_within_at.has_deriv_within_at {f' : 𝕜 →L[𝕜] F} : - has_fderiv_within_at f f' s x → has_deriv_within_at f (f' 1) s x := -has_fderiv_within_at_iff_has_deriv_within_at.mp - -lemma has_deriv_within_at.has_fderiv_within_at {f' : F} : - has_deriv_within_at f f' s x → has_fderiv_within_at f (smul_right (1 : 𝕜 →L[𝕜] 𝕜) f') s x := -has_deriv_within_at_iff_has_fderiv_within_at.mp - -/-- Expressing `has_fderiv_at f f' x` in terms of `has_deriv_at` -/ -lemma has_fderiv_at_iff_has_deriv_at {f' : 𝕜 →L[𝕜] F} : - has_fderiv_at f f' x ↔ has_deriv_at f (f' 1) x := -has_fderiv_at_filter_iff_has_deriv_at_filter - -lemma has_fderiv_at.has_deriv_at {f' : 𝕜 →L[𝕜] F} : - has_fderiv_at f f' x → has_deriv_at f (f' 1) x := -has_fderiv_at_iff_has_deriv_at.mp - -lemma has_strict_fderiv_at_iff_has_strict_deriv_at {f' : 𝕜 →L[𝕜] F} : - has_strict_fderiv_at f f' x ↔ has_strict_deriv_at f (f' 1) x := -by simp [has_strict_deriv_at, has_strict_fderiv_at] - -protected lemma has_strict_fderiv_at.has_strict_deriv_at {f' : 𝕜 →L[𝕜] F} : - has_strict_fderiv_at f f' x → has_strict_deriv_at f (f' 1) x := -has_strict_fderiv_at_iff_has_strict_deriv_at.mp - -lemma has_strict_deriv_at_iff_has_strict_fderiv_at : - has_strict_deriv_at f f' x ↔ has_strict_fderiv_at f (smul_right (1 : 𝕜 →L[𝕜] 𝕜) f') x := -iff.rfl - -alias has_strict_deriv_at_iff_has_strict_fderiv_at ↔ has_strict_deriv_at.has_strict_fderiv_at _ - -/-- Expressing `has_deriv_at f f' x` in terms of `has_fderiv_at` -/ -lemma has_deriv_at_iff_has_fderiv_at {f' : F} : - has_deriv_at f f' x ↔ - has_fderiv_at f (smul_right (1 : 𝕜 →L[𝕜] 𝕜) f') x := -iff.rfl - -alias has_deriv_at_iff_has_fderiv_at ↔ has_deriv_at.has_fderiv_at _ - -lemma deriv_within_zero_of_not_differentiable_within_at - (h : ¬ differentiable_within_at 𝕜 f s x) : deriv_within f s x = 0 := -by { unfold deriv_within, rw fderiv_within_zero_of_not_differentiable_within_at, simp, assumption } - -lemma differentiable_within_at_of_deriv_within_ne_zero (h : deriv_within f s x ≠ 0) : - differentiable_within_at 𝕜 f s x := -not_imp_comm.1 deriv_within_zero_of_not_differentiable_within_at h - -lemma deriv_zero_of_not_differentiable_at (h : ¬ differentiable_at 𝕜 f x) : deriv f x = 0 := -by { unfold deriv, rw fderiv_zero_of_not_differentiable_at, simp, assumption } - -lemma differentiable_at_of_deriv_ne_zero (h : deriv f x ≠ 0) : differentiable_at 𝕜 f x := -not_imp_comm.1 deriv_zero_of_not_differentiable_at h - -theorem unique_diff_within_at.eq_deriv (s : set 𝕜) (H : unique_diff_within_at 𝕜 s x) - (h : has_deriv_within_at f f' s x) (h₁ : has_deriv_within_at f f₁' s x) : f' = f₁' := -smul_right_one_eq_iff.mp $ unique_diff_within_at.eq H h h₁ - -theorem has_deriv_at_filter_iff_is_o : - has_deriv_at_filter f f' x L ↔ (λ x' : 𝕜, f x' - f x - (x' - x) • f') =o[L] (λ x', x' - x) := -iff.rfl - -theorem has_deriv_at_filter_iff_tendsto : - has_deriv_at_filter f f' x L ↔ - tendsto (λ x' : 𝕜, ‖x' - x‖⁻¹ * ‖f x' - f x - (x' - x) • f'‖) L (𝓝 0) := -has_fderiv_at_filter_iff_tendsto - -theorem has_deriv_within_at_iff_is_o : - has_deriv_within_at f f' s x - ↔ (λ x' : 𝕜, f x' - f x - (x' - x) • f') =o[𝓝[s] x] (λ x', x' - x) := -iff.rfl - -theorem has_deriv_within_at_iff_tendsto : has_deriv_within_at f f' s x ↔ - tendsto (λ x', ‖x' - x‖⁻¹ * ‖f x' - f x - (x' - x) • f'‖) (𝓝[s] x) (𝓝 0) := -has_fderiv_at_filter_iff_tendsto - -theorem has_deriv_at_iff_is_o : - has_deriv_at f f' x ↔ (λ x' : 𝕜, f x' - f x - (x' - x) • f') =o[𝓝 x] (λ x', x' - x) := -iff.rfl - -theorem has_deriv_at_iff_tendsto : has_deriv_at f f' x ↔ - tendsto (λ x', ‖x' - x‖⁻¹ * ‖f x' - f x - (x' - x) • f'‖) (𝓝 x) (𝓝 0) := -has_fderiv_at_filter_iff_tendsto - -theorem has_strict_deriv_at.has_deriv_at (h : has_strict_deriv_at f f' x) : - has_deriv_at f f' x := -h.has_fderiv_at - -/-- If the domain has dimension one, then Fréchet derivative is equivalent to the classical -definition with a limit. In this version we have to take the limit along the subset `-{x}`, -because for `y=x` the slope equals zero due to the convention `0⁻¹=0`. -/ -lemma has_deriv_at_filter_iff_tendsto_slope {x : 𝕜} {L : filter 𝕜} : - has_deriv_at_filter f f' x L ↔ tendsto (slope f x) (L ⊓ 𝓟 {x}ᶜ) (𝓝 f') := -begin - conv_lhs { simp only [has_deriv_at_filter_iff_tendsto, (norm_inv _).symm, - (norm_smul _ _).symm, tendsto_zero_iff_norm_tendsto_zero.symm] }, - conv_rhs { rw [← nhds_translation_sub f', tendsto_comap_iff] }, - refine (tendsto_inf_principal_nhds_iff_of_forall_eq $ by simp).symm.trans (tendsto_congr' _), - refine (eventually_principal.2 $ λ z hz, _).filter_mono inf_le_right, - simp only [(∘)], - rw [smul_sub, ← mul_smul, inv_mul_cancel (sub_ne_zero.2 hz), one_smul, slope_def_module] -end - -lemma has_deriv_within_at_iff_tendsto_slope : - has_deriv_within_at f f' s x ↔ tendsto (slope f x) (𝓝[s \ {x}] x) (𝓝 f') := -begin - simp only [has_deriv_within_at, nhds_within, diff_eq, inf_assoc.symm, inf_principal.symm], - exact has_deriv_at_filter_iff_tendsto_slope -end - -lemma has_deriv_within_at_iff_tendsto_slope' (hs : x ∉ s) : - has_deriv_within_at f f' s x ↔ tendsto (slope f x) (𝓝[s] x) (𝓝 f') := -begin - convert ← has_deriv_within_at_iff_tendsto_slope, - exact diff_singleton_eq_self hs -end - -lemma has_deriv_at_iff_tendsto_slope : - has_deriv_at f f' x ↔ tendsto (slope f x) (𝓝[≠] x) (𝓝 f') := -has_deriv_at_filter_iff_tendsto_slope - -theorem has_deriv_within_at_congr_set' {s t : set 𝕜} (y : 𝕜) (h : s =ᶠ[𝓝[{y}ᶜ] x] t) : - has_deriv_within_at f f' s x ↔ has_deriv_within_at f f' t x := -has_fderiv_within_at_congr_set' y h - -theorem has_deriv_within_at_congr_set {s t : set 𝕜} (h : s =ᶠ[𝓝 x] t) : - has_deriv_within_at f f' s x ↔ has_deriv_within_at f f' t x := -has_fderiv_within_at_congr_set h - -alias has_deriv_within_at_congr_set ↔ has_deriv_within_at.congr_set _ - -@[simp] lemma has_deriv_within_at_diff_singleton : - has_deriv_within_at f f' (s \ {x}) x ↔ has_deriv_within_at f f' s x := -has_fderiv_within_at_diff_singleton _ - -@[simp] lemma has_deriv_within_at_Ioi_iff_Ici [partial_order 𝕜] : - has_deriv_within_at f f' (Ioi x) x ↔ has_deriv_within_at f f' (Ici x) x := -by rw [← Ici_diff_left, has_deriv_within_at_diff_singleton] - -alias has_deriv_within_at_Ioi_iff_Ici ↔ - has_deriv_within_at.Ici_of_Ioi has_deriv_within_at.Ioi_of_Ici - -@[simp] lemma has_deriv_within_at_Iio_iff_Iic [partial_order 𝕜] : - has_deriv_within_at f f' (Iio x) x ↔ has_deriv_within_at f f' (Iic x) x := -by rw [← Iic_diff_right, has_deriv_within_at_diff_singleton] - -alias has_deriv_within_at_Iio_iff_Iic ↔ - has_deriv_within_at.Iic_of_Iio has_deriv_within_at.Iio_of_Iic - -theorem has_deriv_within_at.Ioi_iff_Ioo [linear_order 𝕜] [order_closed_topology 𝕜] {x y : 𝕜} - (h : x < y) : - has_deriv_within_at f f' (Ioo x y) x ↔ has_deriv_within_at f f' (Ioi x) x := -has_fderiv_within_at_inter $ Iio_mem_nhds h - -alias has_deriv_within_at.Ioi_iff_Ioo ↔ - has_deriv_within_at.Ioi_of_Ioo has_deriv_within_at.Ioo_of_Ioi - -theorem has_deriv_at_iff_is_o_nhds_zero : has_deriv_at f f' x ↔ - (λh, f (x + h) - f x - h • f') =o[𝓝 0] (λh, h) := -has_fderiv_at_iff_is_o_nhds_zero - -theorem has_deriv_at_filter.mono (h : has_deriv_at_filter f f' x L₂) (hst : L₁ ≤ L₂) : - has_deriv_at_filter f f' x L₁ := -has_fderiv_at_filter.mono h hst - -theorem has_deriv_within_at.mono (h : has_deriv_within_at f f' t x) (hst : s ⊆ t) : - has_deriv_within_at f f' s x := -has_fderiv_within_at.mono h hst - -theorem has_deriv_within_at.mono_of_mem (h : has_deriv_within_at f f' t x) (hst : t ∈ 𝓝[s] x) : - has_deriv_within_at f f' s x := -has_fderiv_within_at.mono_of_mem h hst - -theorem has_deriv_at.has_deriv_at_filter (h : has_deriv_at f f' x) (hL : L ≤ 𝓝 x) : - has_deriv_at_filter f f' x L := -has_fderiv_at.has_fderiv_at_filter h hL - -theorem has_deriv_at.has_deriv_within_at - (h : has_deriv_at f f' x) : has_deriv_within_at f f' s x := -has_fderiv_at.has_fderiv_within_at h - -lemma has_deriv_within_at.differentiable_within_at (h : has_deriv_within_at f f' s x) : - differentiable_within_at 𝕜 f s x := -has_fderiv_within_at.differentiable_within_at h - -lemma has_deriv_at.differentiable_at (h : has_deriv_at f f' x) : differentiable_at 𝕜 f x := -has_fderiv_at.differentiable_at h - -@[simp] lemma has_deriv_within_at_univ : has_deriv_within_at f f' univ x ↔ has_deriv_at f f' x := -has_fderiv_within_at_univ - -theorem has_deriv_at.unique - (h₀ : has_deriv_at f f₀' x) (h₁ : has_deriv_at f f₁' x) : f₀' = f₁' := -smul_right_one_eq_iff.mp $ h₀.has_fderiv_at.unique h₁ - -lemma has_deriv_within_at_inter' (h : t ∈ 𝓝[s] x) : - has_deriv_within_at f f' (s ∩ t) x ↔ has_deriv_within_at f f' s x := -has_fderiv_within_at_inter' h - -lemma has_deriv_within_at_inter (h : t ∈ 𝓝 x) : - has_deriv_within_at f f' (s ∩ t) x ↔ has_deriv_within_at f f' s x := -has_fderiv_within_at_inter h - -lemma has_deriv_within_at.union (hs : has_deriv_within_at f f' s x) - (ht : has_deriv_within_at f f' t x) : - has_deriv_within_at f f' (s ∪ t) x := -hs.has_fderiv_within_at.union ht.has_fderiv_within_at - -lemma has_deriv_within_at.nhds_within (h : has_deriv_within_at f f' s x) - (ht : s ∈ 𝓝[t] x) : has_deriv_within_at f f' t x := -(has_deriv_within_at_inter' ht).1 (h.mono (inter_subset_right _ _)) - -lemma has_deriv_within_at.has_deriv_at (h : has_deriv_within_at f f' s x) (hs : s ∈ 𝓝 x) : - has_deriv_at f f' x := -has_fderiv_within_at.has_fderiv_at h hs - -lemma differentiable_within_at.has_deriv_within_at (h : differentiable_within_at 𝕜 f s x) : - has_deriv_within_at f (deriv_within f s x) s x := -h.has_fderiv_within_at.has_deriv_within_at - -lemma differentiable_at.has_deriv_at (h : differentiable_at 𝕜 f x) : has_deriv_at f (deriv f x) x := -h.has_fderiv_at.has_deriv_at - -@[simp] lemma has_deriv_at_deriv_iff : has_deriv_at f (deriv f x) x ↔ differentiable_at 𝕜 f x := -⟨λ h, h.differentiable_at, λ h, h.has_deriv_at⟩ - -@[simp] lemma has_deriv_within_at_deriv_within_iff : - has_deriv_within_at f (deriv_within f s x) s x ↔ differentiable_within_at 𝕜 f s x := -⟨λ h, h.differentiable_within_at, λ h, h.has_deriv_within_at⟩ - -lemma differentiable_on.has_deriv_at (h : differentiable_on 𝕜 f s) (hs : s ∈ 𝓝 x) : - has_deriv_at f (deriv f x) x := -(h.has_fderiv_at hs).has_deriv_at - -lemma has_deriv_at.deriv (h : has_deriv_at f f' x) : deriv f x = f' := -h.differentiable_at.has_deriv_at.unique h - -lemma deriv_eq {f' : 𝕜 → F} (h : ∀ x, has_deriv_at f (f' x) x) : deriv f = f' := -funext $ λ x, (h x).deriv - -lemma has_deriv_within_at.deriv_within - (h : has_deriv_within_at f f' s x) (hxs : unique_diff_within_at 𝕜 s x) : - deriv_within f s x = f' := -hxs.eq_deriv _ h.differentiable_within_at.has_deriv_within_at h - -lemma fderiv_within_deriv_within : (fderiv_within 𝕜 f s x : 𝕜 → F) 1 = deriv_within f s x := -rfl - -lemma deriv_within_fderiv_within : - smul_right (1 : 𝕜 →L[𝕜] 𝕜) (deriv_within f s x) = fderiv_within 𝕜 f s x := -by simp [deriv_within] - -lemma fderiv_deriv : (fderiv 𝕜 f x : 𝕜 → F) 1 = deriv f x := -rfl - -lemma deriv_fderiv : - smul_right (1 : 𝕜 →L[𝕜] 𝕜) (deriv f x) = fderiv 𝕜 f x := -by simp [deriv] - -lemma differentiable_at.deriv_within (h : differentiable_at 𝕜 f x) - (hxs : unique_diff_within_at 𝕜 s x) : deriv_within f s x = deriv f x := -by { unfold deriv_within deriv, rw h.fderiv_within hxs } - -theorem has_deriv_within_at.deriv_eq_zero (hd : has_deriv_within_at f 0 s x) - (H : unique_diff_within_at 𝕜 s x) : deriv f x = 0 := -(em' (differentiable_at 𝕜 f x)).elim deriv_zero_of_not_differentiable_at $ - λ h, H.eq_deriv _ h.has_deriv_at.has_deriv_within_at hd - -lemma deriv_within_of_mem (st : t ∈ 𝓝[s] x) (ht : unique_diff_within_at 𝕜 s x) - (h : differentiable_within_at 𝕜 f t x) : - deriv_within f s x = deriv_within f t x := -((differentiable_within_at.has_deriv_within_at h).mono_of_mem st).deriv_within ht - -lemma deriv_within_subset (st : s ⊆ t) (ht : unique_diff_within_at 𝕜 s x) - (h : differentiable_within_at 𝕜 f t x) : - deriv_within f s x = deriv_within f t x := -((differentiable_within_at.has_deriv_within_at h).mono st).deriv_within ht - -lemma deriv_within_congr_set' (y : 𝕜) (h : s =ᶠ[𝓝[{y}ᶜ] x] t) : - deriv_within f s x = deriv_within f t x := -by simp only [deriv_within, fderiv_within_congr_set' y h] - -lemma deriv_within_congr_set (h : s =ᶠ[𝓝 x] t) : deriv_within f s x = deriv_within f t x := -by simp only [deriv_within, fderiv_within_congr_set h] - -@[simp] lemma deriv_within_univ : deriv_within f univ = deriv f := -by { ext, unfold deriv_within deriv, rw fderiv_within_univ } - -lemma deriv_within_inter (ht : t ∈ 𝓝 x) : - deriv_within f (s ∩ t) x = deriv_within f s x := -by { unfold deriv_within, rw fderiv_within_inter ht } - -lemma deriv_within_of_open (hs : is_open s) (hx : x ∈ s) : - deriv_within f s x = deriv f x := -by { unfold deriv_within, rw fderiv_within_of_open hs hx, refl } - -lemma deriv_mem_iff {f : 𝕜 → F} {s : set F} {x : 𝕜} : - deriv f x ∈ s ↔ (differentiable_at 𝕜 f x ∧ deriv f x ∈ s) ∨ - (¬differentiable_at 𝕜 f x ∧ (0 : F) ∈ s) := -by by_cases hx : differentiable_at 𝕜 f x; simp [deriv_zero_of_not_differentiable_at, *] - -lemma deriv_within_mem_iff {f : 𝕜 → F} {t : set 𝕜} {s : set F} {x : 𝕜} : - deriv_within f t x ∈ s ↔ (differentiable_within_at 𝕜 f t x ∧ deriv_within f t x ∈ s) ∨ - (¬differentiable_within_at 𝕜 f t x ∧ (0 : F) ∈ s) := -by by_cases hx : differentiable_within_at 𝕜 f t x; - simp [deriv_within_zero_of_not_differentiable_within_at, *] - -lemma differentiable_within_at_Ioi_iff_Ici [partial_order 𝕜] : - differentiable_within_at 𝕜 f (Ioi x) x ↔ differentiable_within_at 𝕜 f (Ici x) x := -⟨λ h, h.has_deriv_within_at.Ici_of_Ioi.differentiable_within_at, -λ h, h.has_deriv_within_at.Ioi_of_Ici.differentiable_within_at⟩ - -lemma deriv_within_Ioi_eq_Ici {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] (f : ℝ → E) - (x : ℝ) : - deriv_within f (Ioi x) x = deriv_within f (Ici x) x := -begin - by_cases H : differentiable_within_at ℝ f (Ioi x) x, - { have A := H.has_deriv_within_at.Ici_of_Ioi, - have B := (differentiable_within_at_Ioi_iff_Ici.1 H).has_deriv_within_at, - simpa using (unique_diff_on_Ici x).eq le_rfl A B }, - { rw [deriv_within_zero_of_not_differentiable_within_at H, - deriv_within_zero_of_not_differentiable_within_at], - rwa differentiable_within_at_Ioi_iff_Ici at H } -end - -section congr -/-! ### Congruence properties of derivatives -/ - -theorem filter.eventually_eq.has_deriv_at_filter_iff - (h₀ : f₀ =ᶠ[L] f₁) (hx : f₀ x = f₁ x) (h₁ : f₀' = f₁') : - has_deriv_at_filter f₀ f₀' x L ↔ has_deriv_at_filter f₁ f₁' x L := -h₀.has_fderiv_at_filter_iff hx (by simp [h₁]) - -lemma has_deriv_at_filter.congr_of_eventually_eq (h : has_deriv_at_filter f f' x L) - (hL : f₁ =ᶠ[L] f) (hx : f₁ x = f x) : has_deriv_at_filter f₁ f' x L := -by rwa hL.has_deriv_at_filter_iff hx rfl - -lemma has_deriv_within_at.congr_mono (h : has_deriv_within_at f f' s x) (ht : ∀x ∈ t, f₁ x = f x) - (hx : f₁ x = f x) (h₁ : t ⊆ s) : has_deriv_within_at f₁ f' t x := -has_fderiv_within_at.congr_mono h ht hx h₁ - -lemma has_deriv_within_at.congr (h : has_deriv_within_at f f' s x) (hs : ∀x ∈ s, f₁ x = f x) - (hx : f₁ x = f x) : has_deriv_within_at f₁ f' s x := -h.congr_mono hs hx (subset.refl _) - -lemma has_deriv_within_at.congr_of_mem (h : has_deriv_within_at f f' s x) (hs : ∀x ∈ s, f₁ x = f x) - (hx : x ∈ s) : has_deriv_within_at f₁ f' s x := -h.congr hs (hs _ hx) - -lemma has_deriv_within_at.congr_of_eventually_eq (h : has_deriv_within_at f f' s x) - (h₁ : f₁ =ᶠ[𝓝[s] x] f) (hx : f₁ x = f x) : has_deriv_within_at f₁ f' s x := -has_deriv_at_filter.congr_of_eventually_eq h h₁ hx - -lemma has_deriv_within_at.congr_of_eventually_eq_of_mem (h : has_deriv_within_at f f' s x) - (h₁ : f₁ =ᶠ[𝓝[s] x] f) (hx : x ∈ s) : has_deriv_within_at f₁ f' s x := -h.congr_of_eventually_eq h₁ (h₁.eq_of_nhds_within hx) - -lemma has_deriv_at.congr_of_eventually_eq (h : has_deriv_at f f' x) - (h₁ : f₁ =ᶠ[𝓝 x] f) : has_deriv_at f₁ f' x := -has_deriv_at_filter.congr_of_eventually_eq h h₁ (mem_of_mem_nhds h₁ : _) - -lemma filter.eventually_eq.deriv_within_eq (hL : f₁ =ᶠ[𝓝[s] x] f) (hx : f₁ x = f x) : - deriv_within f₁ s x = deriv_within f s x := -by { unfold deriv_within, rw hL.fderiv_within_eq hx } - -lemma deriv_within_congr (hs : eq_on f₁ f s) (hx : f₁ x = f x) : - deriv_within f₁ s x = deriv_within f s x := -by { unfold deriv_within, rw fderiv_within_congr hs hx } - -lemma filter.eventually_eq.deriv_eq (hL : f₁ =ᶠ[𝓝 x] f) : deriv f₁ x = deriv f x := -by { unfold deriv, rwa filter.eventually_eq.fderiv_eq } - -protected lemma filter.eventually_eq.deriv (h : f₁ =ᶠ[𝓝 x] f) : deriv f₁ =ᶠ[𝓝 x] deriv f := -h.eventually_eq_nhds.mono $ λ x h, h.deriv_eq - -end congr - -section id -/-! ### Derivative of the identity -/ -variables (s x L) - -theorem has_deriv_at_filter_id : has_deriv_at_filter id 1 x L := -(has_fderiv_at_filter_id x L).has_deriv_at_filter - -theorem has_deriv_within_at_id : has_deriv_within_at id 1 s x := -has_deriv_at_filter_id _ _ - -theorem has_deriv_at_id : has_deriv_at id 1 x := -has_deriv_at_filter_id _ _ - -theorem has_deriv_at_id' : has_deriv_at (λ (x : 𝕜), x) 1 x := -has_deriv_at_filter_id _ _ - -theorem has_strict_deriv_at_id : has_strict_deriv_at id 1 x := -(has_strict_fderiv_at_id x).has_strict_deriv_at - -lemma deriv_id : deriv id x = 1 := -has_deriv_at.deriv (has_deriv_at_id x) - -@[simp] lemma deriv_id' : deriv (@id 𝕜) = λ _, 1 := funext deriv_id - -@[simp] lemma deriv_id'' : deriv (λ x : 𝕜, x) = λ _, 1 := deriv_id' - -lemma deriv_within_id (hxs : unique_diff_within_at 𝕜 s x) : deriv_within id s x = 1 := -(has_deriv_within_at_id x s).deriv_within hxs - -end id - -section const -/-! ### Derivative of constant functions -/ -variables (c : F) (s x L) - -theorem has_deriv_at_filter_const : has_deriv_at_filter (λ x, c) 0 x L := -(has_fderiv_at_filter_const c x L).has_deriv_at_filter - -theorem has_strict_deriv_at_const : has_strict_deriv_at (λ x, c) 0 x := -(has_strict_fderiv_at_const c x).has_strict_deriv_at - -theorem has_deriv_within_at_const : has_deriv_within_at (λ x, c) 0 s x := -has_deriv_at_filter_const _ _ _ - -theorem has_deriv_at_const : has_deriv_at (λ x, c) 0 x := -has_deriv_at_filter_const _ _ _ - -lemma deriv_const : deriv (λ x, c) x = 0 := -has_deriv_at.deriv (has_deriv_at_const x c) - -@[simp] lemma deriv_const' : deriv (λ x:𝕜, c) = λ x, 0 := -funext (λ x, deriv_const x c) - -lemma deriv_within_const (hxs : unique_diff_within_at 𝕜 s x) : deriv_within (λ x, c) s x = 0 := -(has_deriv_within_at_const _ _ _).deriv_within hxs - -end const - -section continuous_linear_map -/-! ### Derivative of continuous linear maps -/ -variables (e : 𝕜 →L[𝕜] F) - -protected lemma continuous_linear_map.has_deriv_at_filter : has_deriv_at_filter e (e 1) x L := -e.has_fderiv_at_filter.has_deriv_at_filter - -protected lemma continuous_linear_map.has_strict_deriv_at : has_strict_deriv_at e (e 1) x := -e.has_strict_fderiv_at.has_strict_deriv_at - -protected lemma continuous_linear_map.has_deriv_at : has_deriv_at e (e 1) x := -e.has_deriv_at_filter - -protected lemma continuous_linear_map.has_deriv_within_at : has_deriv_within_at e (e 1) s x := -e.has_deriv_at_filter - -@[simp] protected lemma continuous_linear_map.deriv : deriv e x = e 1 := -e.has_deriv_at.deriv - -protected lemma continuous_linear_map.deriv_within (hxs : unique_diff_within_at 𝕜 s x) : - deriv_within e s x = e 1 := -e.has_deriv_within_at.deriv_within hxs - -end continuous_linear_map - -section linear_map -/-! ### Derivative of bundled linear maps -/ -variables (e : 𝕜 →ₗ[𝕜] F) - -protected lemma linear_map.has_deriv_at_filter : has_deriv_at_filter e (e 1) x L := -e.to_continuous_linear_map₁.has_deriv_at_filter - -protected lemma linear_map.has_strict_deriv_at : has_strict_deriv_at e (e 1) x := -e.to_continuous_linear_map₁.has_strict_deriv_at - -protected lemma linear_map.has_deriv_at : has_deriv_at e (e 1) x := -e.has_deriv_at_filter - -protected lemma linear_map.has_deriv_within_at : has_deriv_within_at e (e 1) s x := -e.has_deriv_at_filter - -@[simp] protected lemma linear_map.deriv : deriv e x = e 1 := -e.has_deriv_at.deriv - -protected lemma linear_map.deriv_within (hxs : unique_diff_within_at 𝕜 s x) : - deriv_within e s x = e 1 := -e.has_deriv_within_at.deriv_within hxs - -end linear_map - -section add -/-! ### Derivative of the sum of two functions -/ - -theorem has_deriv_at_filter.add - (hf : has_deriv_at_filter f f' x L) (hg : has_deriv_at_filter g g' x L) : - has_deriv_at_filter (λ y, f y + g y) (f' + g') x L := -by simpa using (hf.add hg).has_deriv_at_filter - -theorem has_strict_deriv_at.add - (hf : has_strict_deriv_at f f' x) (hg : has_strict_deriv_at g g' x) : - has_strict_deriv_at (λ y, f y + g y) (f' + g') x := -by simpa using (hf.add hg).has_strict_deriv_at - -theorem has_deriv_within_at.add - (hf : has_deriv_within_at f f' s x) (hg : has_deriv_within_at g g' s x) : - has_deriv_within_at (λ y, f y + g y) (f' + g') s x := -hf.add hg - -theorem has_deriv_at.add - (hf : has_deriv_at f f' x) (hg : has_deriv_at g g' x) : - has_deriv_at (λ x, f x + g x) (f' + g') x := -hf.add hg - -lemma deriv_within_add (hxs : unique_diff_within_at 𝕜 s x) - (hf : differentiable_within_at 𝕜 f s x) (hg : differentiable_within_at 𝕜 g s x) : - deriv_within (λy, f y + g y) s x = deriv_within f s x + deriv_within g s x := -(hf.has_deriv_within_at.add hg.has_deriv_within_at).deriv_within hxs - -@[simp] lemma deriv_add - (hf : differentiable_at 𝕜 f x) (hg : differentiable_at 𝕜 g x) : - deriv (λy, f y + g y) x = deriv f x + deriv g x := -(hf.has_deriv_at.add hg.has_deriv_at).deriv - -theorem has_deriv_at_filter.add_const - (hf : has_deriv_at_filter f f' x L) (c : F) : - has_deriv_at_filter (λ y, f y + c) f' x L := -add_zero f' ▸ hf.add (has_deriv_at_filter_const x L c) - -theorem has_deriv_within_at.add_const - (hf : has_deriv_within_at f f' s x) (c : F) : - has_deriv_within_at (λ y, f y + c) f' s x := -hf.add_const c - -theorem has_deriv_at.add_const - (hf : has_deriv_at f f' x) (c : F) : - has_deriv_at (λ x, f x + c) f' x := -hf.add_const c - -lemma deriv_within_add_const (hxs : unique_diff_within_at 𝕜 s x) (c : F) : - deriv_within (λy, f y + c) s x = deriv_within f s x := -by simp only [deriv_within, fderiv_within_add_const hxs] - -lemma deriv_add_const (c : F) : deriv (λy, f y + c) x = deriv f x := -by simp only [deriv, fderiv_add_const] - -@[simp] lemma deriv_add_const' (c : F) : deriv (λ y, f y + c) = deriv f := -funext $ λ x, deriv_add_const c - -theorem has_deriv_at_filter.const_add (c : F) (hf : has_deriv_at_filter f f' x L) : - has_deriv_at_filter (λ y, c + f y) f' x L := -zero_add f' ▸ (has_deriv_at_filter_const x L c).add hf - -theorem has_deriv_within_at.const_add (c : F) (hf : has_deriv_within_at f f' s x) : - has_deriv_within_at (λ y, c + f y) f' s x := -hf.const_add c - -theorem has_deriv_at.const_add (c : F) (hf : has_deriv_at f f' x) : - has_deriv_at (λ x, c + f x) f' x := -hf.const_add c - -lemma deriv_within_const_add (hxs : unique_diff_within_at 𝕜 s x) (c : F) : - deriv_within (λy, c + f y) s x = deriv_within f s x := -by simp only [deriv_within, fderiv_within_const_add hxs] - -lemma deriv_const_add (c : F) : deriv (λy, c + f y) x = deriv f x := -by simp only [deriv, fderiv_const_add] - -@[simp] lemma deriv_const_add' (c : F) : deriv (λ y, c + f y) = deriv f := -funext $ λ x, deriv_const_add c - -end add - -section sum -/-! ### Derivative of a finite sum of functions -/ - -open_locale big_operators - -variables {ι : Type*} {u : finset ι} {A : ι → (𝕜 → F)} {A' : ι → F} - -theorem has_deriv_at_filter.sum (h : ∀ i ∈ u, has_deriv_at_filter (A i) (A' i) x L) : - has_deriv_at_filter (λ y, ∑ i in u, A i y) (∑ i in u, A' i) x L := -by simpa [continuous_linear_map.sum_apply] using (has_fderiv_at_filter.sum h).has_deriv_at_filter - -theorem has_strict_deriv_at.sum (h : ∀ i ∈ u, has_strict_deriv_at (A i) (A' i) x) : - has_strict_deriv_at (λ y, ∑ i in u, A i y) (∑ i in u, A' i) x := -by simpa [continuous_linear_map.sum_apply] using (has_strict_fderiv_at.sum h).has_strict_deriv_at - -theorem has_deriv_within_at.sum (h : ∀ i ∈ u, has_deriv_within_at (A i) (A' i) s x) : - has_deriv_within_at (λ y, ∑ i in u, A i y) (∑ i in u, A' i) s x := -has_deriv_at_filter.sum h - -theorem has_deriv_at.sum (h : ∀ i ∈ u, has_deriv_at (A i) (A' i) x) : - has_deriv_at (λ y, ∑ i in u, A i y) (∑ i in u, A' i) x := -has_deriv_at_filter.sum h - -lemma deriv_within_sum (hxs : unique_diff_within_at 𝕜 s x) - (h : ∀ i ∈ u, differentiable_within_at 𝕜 (A i) s x) : - deriv_within (λ y, ∑ i in u, A i y) s x = ∑ i in u, deriv_within (A i) s x := -(has_deriv_within_at.sum (λ i hi, (h i hi).has_deriv_within_at)).deriv_within hxs - -@[simp] lemma deriv_sum (h : ∀ i ∈ u, differentiable_at 𝕜 (A i) x) : - deriv (λ y, ∑ i in u, A i y) x = ∑ i in u, deriv (A i) x := -(has_deriv_at.sum (λ i hi, (h i hi).has_deriv_at)).deriv - -end sum - -section pi - -/-! ### Derivatives of functions `f : 𝕜 → Π i, E i` -/ - -variables {ι : Type*} [fintype ι] {E' : ι → Type*} [Π i, normed_add_comm_group (E' i)] - [Π i, normed_space 𝕜 (E' i)] {φ : 𝕜 → Π i, E' i} {φ' : Π i, E' i} - -@[simp] lemma has_strict_deriv_at_pi : - has_strict_deriv_at φ φ' x ↔ ∀ i, has_strict_deriv_at (λ x, φ x i) (φ' i) x := -has_strict_fderiv_at_pi' - -@[simp] lemma has_deriv_at_filter_pi : - has_deriv_at_filter φ φ' x L ↔ - ∀ i, has_deriv_at_filter (λ x, φ x i) (φ' i) x L := -has_fderiv_at_filter_pi' - -lemma has_deriv_at_pi : - has_deriv_at φ φ' x ↔ ∀ i, has_deriv_at (λ x, φ x i) (φ' i) x:= -has_deriv_at_filter_pi - -lemma has_deriv_within_at_pi : - has_deriv_within_at φ φ' s x ↔ ∀ i, has_deriv_within_at (λ x, φ x i) (φ' i) s x:= -has_deriv_at_filter_pi - -lemma deriv_within_pi (h : ∀ i, differentiable_within_at 𝕜 (λ x, φ x i) s x) - (hs : unique_diff_within_at 𝕜 s x) : - deriv_within φ s x = λ i, deriv_within (λ x, φ x i) s x := -(has_deriv_within_at_pi.2 (λ i, (h i).has_deriv_within_at)).deriv_within hs - -lemma deriv_pi (h : ∀ i, differentiable_at 𝕜 (λ x, φ x i) x) : - deriv φ x = λ i, deriv (λ x, φ x i) x := -(has_deriv_at_pi.2 (λ i, (h i).has_deriv_at)).deriv - -end pi - -section smul - -/-! ### Derivative of the multiplication of a scalar function and a vector function -/ - -variables {𝕜' : Type*} [nontrivially_normed_field 𝕜'] [normed_algebra 𝕜 𝕜'] - [normed_space 𝕜' F] [is_scalar_tower 𝕜 𝕜' F] {c : 𝕜 → 𝕜'} {c' : 𝕜'} - -theorem has_deriv_within_at.smul - (hc : has_deriv_within_at c c' s x) (hf : has_deriv_within_at f f' s x) : - has_deriv_within_at (λ y, c y • f y) (c x • f' + c' • f x) s x := -by simpa using (has_fderiv_within_at.smul hc hf).has_deriv_within_at - -theorem has_deriv_at.smul - (hc : has_deriv_at c c' x) (hf : has_deriv_at f f' x) : - has_deriv_at (λ y, c y • f y) (c x • f' + c' • f x) x := -begin - rw [← has_deriv_within_at_univ] at *, - exact hc.smul hf -end - -theorem has_strict_deriv_at.smul - (hc : has_strict_deriv_at c c' x) (hf : has_strict_deriv_at f f' x) : - has_strict_deriv_at (λ y, c y • f y) (c x • f' + c' • f x) x := -by simpa using (hc.smul hf).has_strict_deriv_at - -lemma deriv_within_smul (hxs : unique_diff_within_at 𝕜 s x) - (hc : differentiable_within_at 𝕜 c s x) (hf : differentiable_within_at 𝕜 f s x) : - deriv_within (λ y, c y • f y) s x = c x • deriv_within f s x + (deriv_within c s x) • f x := -(hc.has_deriv_within_at.smul hf.has_deriv_within_at).deriv_within hxs - -lemma deriv_smul (hc : differentiable_at 𝕜 c x) (hf : differentiable_at 𝕜 f x) : - deriv (λ y, c y • f y) x = c x • deriv f x + (deriv c x) • f x := -(hc.has_deriv_at.smul hf.has_deriv_at).deriv - -theorem has_strict_deriv_at.smul_const - (hc : has_strict_deriv_at c c' x) (f : F) : - has_strict_deriv_at (λ y, c y • f) (c' • f) x := -begin - have := hc.smul (has_strict_deriv_at_const x f), - rwa [smul_zero, zero_add] at this, -end - -theorem has_deriv_within_at.smul_const - (hc : has_deriv_within_at c c' s x) (f : F) : - has_deriv_within_at (λ y, c y • f) (c' • f) s x := -begin - have := hc.smul (has_deriv_within_at_const x s f), - rwa [smul_zero, zero_add] at this -end - -theorem has_deriv_at.smul_const - (hc : has_deriv_at c c' x) (f : F) : - has_deriv_at (λ y, c y • f) (c' • f) x := -begin - rw [← has_deriv_within_at_univ] at *, - exact hc.smul_const f -end - -lemma deriv_within_smul_const (hxs : unique_diff_within_at 𝕜 s x) - (hc : differentiable_within_at 𝕜 c s x) (f : F) : - deriv_within (λ y, c y • f) s x = (deriv_within c s x) • f := -(hc.has_deriv_within_at.smul_const f).deriv_within hxs - -lemma deriv_smul_const (hc : differentiable_at 𝕜 c x) (f : F) : - deriv (λ y, c y • f) x = (deriv c x) • f := -(hc.has_deriv_at.smul_const f).deriv - -end smul - -section const_smul - -variables {R : Type*} [semiring R] [module R F] [smul_comm_class 𝕜 R F] - [has_continuous_const_smul R F] - -theorem has_strict_deriv_at.const_smul - (c : R) (hf : has_strict_deriv_at f f' x) : - has_strict_deriv_at (λ y, c • f y) (c • f') x := -by simpa using (hf.const_smul c).has_strict_deriv_at - -theorem has_deriv_at_filter.const_smul - (c : R) (hf : has_deriv_at_filter f f' x L) : - has_deriv_at_filter (λ y, c • f y) (c • f') x L := -by simpa using (hf.const_smul c).has_deriv_at_filter - -theorem has_deriv_within_at.const_smul - (c : R) (hf : has_deriv_within_at f f' s x) : - has_deriv_within_at (λ y, c • f y) (c • f') s x := -hf.const_smul c - -theorem has_deriv_at.const_smul (c : R) (hf : has_deriv_at f f' x) : - has_deriv_at (λ y, c • f y) (c • f') x := -hf.const_smul c - -lemma deriv_within_const_smul (hxs : unique_diff_within_at 𝕜 s x) - (c : R) (hf : differentiable_within_at 𝕜 f s x) : - deriv_within (λ y, c • f y) s x = c • deriv_within f s x := -(hf.has_deriv_within_at.const_smul c).deriv_within hxs - -lemma deriv_const_smul (c : R) (hf : differentiable_at 𝕜 f x) : - deriv (λ y, c • f y) x = c • deriv f x := -(hf.has_deriv_at.const_smul c).deriv - -end const_smul - -section neg -/-! ### Derivative of the negative of a function -/ - -theorem has_deriv_at_filter.neg (h : has_deriv_at_filter f f' x L) : - has_deriv_at_filter (λ x, -f x) (-f') x L := -by simpa using h.neg.has_deriv_at_filter - -theorem has_deriv_within_at.neg (h : has_deriv_within_at f f' s x) : - has_deriv_within_at (λ x, -f x) (-f') s x := -h.neg - -theorem has_deriv_at.neg (h : has_deriv_at f f' x) : has_deriv_at (λ x, -f x) (-f') x := -h.neg - -theorem has_strict_deriv_at.neg (h : has_strict_deriv_at f f' x) : - has_strict_deriv_at (λ x, -f x) (-f') x := -by simpa using h.neg.has_strict_deriv_at - -lemma deriv_within.neg (hxs : unique_diff_within_at 𝕜 s x) : - deriv_within (λy, -f y) s x = - deriv_within f s x := -by simp only [deriv_within, fderiv_within_neg hxs, continuous_linear_map.neg_apply] - -lemma deriv.neg : deriv (λy, -f y) x = - deriv f x := -by simp only [deriv, fderiv_neg, continuous_linear_map.neg_apply] - -@[simp] lemma deriv.neg' : deriv (λy, -f y) = (λ x, - deriv f x) := -funext $ λ x, deriv.neg - -end neg - -section neg2 -/-! ### Derivative of the negation function (i.e `has_neg.neg`) -/ - -variables (s x L) - -theorem has_deriv_at_filter_neg : has_deriv_at_filter has_neg.neg (-1) x L := -has_deriv_at_filter.neg $ has_deriv_at_filter_id _ _ - -theorem has_deriv_within_at_neg : has_deriv_within_at has_neg.neg (-1) s x := -has_deriv_at_filter_neg _ _ - -theorem has_deriv_at_neg : has_deriv_at has_neg.neg (-1) x := -has_deriv_at_filter_neg _ _ - -theorem has_deriv_at_neg' : has_deriv_at (λ x, -x) (-1) x := -has_deriv_at_filter_neg _ _ - -theorem has_strict_deriv_at_neg : has_strict_deriv_at has_neg.neg (-1) x := -has_strict_deriv_at.neg $ has_strict_deriv_at_id _ - -lemma deriv_neg : deriv has_neg.neg x = -1 := -has_deriv_at.deriv (has_deriv_at_neg x) - -@[simp] lemma deriv_neg' : deriv (has_neg.neg : 𝕜 → 𝕜) = λ _, -1 := -funext deriv_neg - -@[simp] lemma deriv_neg'' : deriv (λ x : 𝕜, -x) x = -1 := -deriv_neg x - -lemma deriv_within_neg (hxs : unique_diff_within_at 𝕜 s x) : deriv_within has_neg.neg s x = -1 := -(has_deriv_within_at_neg x s).deriv_within hxs - -lemma differentiable_neg : differentiable 𝕜 (has_neg.neg : 𝕜 → 𝕜) := -differentiable.neg differentiable_id - -lemma differentiable_on_neg : differentiable_on 𝕜 (has_neg.neg : 𝕜 → 𝕜) s := -differentiable_on.neg differentiable_on_id - -end neg2 - -section sub -/-! ### Derivative of the difference of two functions -/ - -theorem has_deriv_at_filter.sub - (hf : has_deriv_at_filter f f' x L) (hg : has_deriv_at_filter g g' x L) : - has_deriv_at_filter (λ x, f x - g x) (f' - g') x L := -by simpa only [sub_eq_add_neg] using hf.add hg.neg - -theorem has_deriv_within_at.sub - (hf : has_deriv_within_at f f' s x) (hg : has_deriv_within_at g g' s x) : - has_deriv_within_at (λ x, f x - g x) (f' - g') s x := -hf.sub hg - -theorem has_deriv_at.sub - (hf : has_deriv_at f f' x) (hg : has_deriv_at g g' x) : - has_deriv_at (λ x, f x - g x) (f' - g') x := -hf.sub hg - -theorem has_strict_deriv_at.sub - (hf : has_strict_deriv_at f f' x) (hg : has_strict_deriv_at g g' x) : - has_strict_deriv_at (λ x, f x - g x) (f' - g') x := -by simpa only [sub_eq_add_neg] using hf.add hg.neg - -lemma deriv_within_sub (hxs : unique_diff_within_at 𝕜 s x) - (hf : differentiable_within_at 𝕜 f s x) (hg : differentiable_within_at 𝕜 g s x) : - deriv_within (λy, f y - g y) s x = deriv_within f s x - deriv_within g s x := -(hf.has_deriv_within_at.sub hg.has_deriv_within_at).deriv_within hxs - -@[simp] lemma deriv_sub - (hf : differentiable_at 𝕜 f x) (hg : differentiable_at 𝕜 g x) : - deriv (λ y, f y - g y) x = deriv f x - deriv g x := -(hf.has_deriv_at.sub hg.has_deriv_at).deriv - -theorem has_deriv_at_filter.is_O_sub (h : has_deriv_at_filter f f' x L) : - (λ x', f x' - f x) =O[L] (λ x', x' - x) := -has_fderiv_at_filter.is_O_sub h - -theorem has_deriv_at_filter.is_O_sub_rev (hf : has_deriv_at_filter f f' x L) (hf' : f' ≠ 0) : - (λ x', x' - x) =O[L] (λ x', f x' - f x) := -suffices antilipschitz_with ‖f'‖₊⁻¹ (smul_right (1 : 𝕜 →L[𝕜] 𝕜) f'), from hf.is_O_sub_rev this, -add_monoid_hom_class.antilipschitz_of_bound (smul_right (1 : 𝕜 →L[𝕜] 𝕜) f') $ - λ x, by simp [norm_smul, ← div_eq_inv_mul, mul_div_cancel _ (mt norm_eq_zero.1 hf')] - -theorem has_deriv_at_filter.sub_const - (hf : has_deriv_at_filter f f' x L) (c : F) : - has_deriv_at_filter (λ x, f x - c) f' x L := -by simpa only [sub_eq_add_neg] using hf.add_const (-c) - -theorem has_deriv_within_at.sub_const - (hf : has_deriv_within_at f f' s x) (c : F) : - has_deriv_within_at (λ x, f x - c) f' s x := -hf.sub_const c - -theorem has_deriv_at.sub_const - (hf : has_deriv_at f f' x) (c : F) : - has_deriv_at (λ x, f x - c) f' x := -hf.sub_const c - -lemma deriv_within_sub_const (hxs : unique_diff_within_at 𝕜 s x) (c : F) : - deriv_within (λy, f y - c) s x = deriv_within f s x := -by simp only [deriv_within, fderiv_within_sub_const hxs] - -lemma deriv_sub_const (c : F) : deriv (λ y, f y - c) x = deriv f x := -by simp only [deriv, fderiv_sub_const] - -theorem has_deriv_at_filter.const_sub (c : F) (hf : has_deriv_at_filter f f' x L) : - has_deriv_at_filter (λ x, c - f x) (-f') x L := -by simpa only [sub_eq_add_neg] using hf.neg.const_add c - -theorem has_deriv_within_at.const_sub (c : F) (hf : has_deriv_within_at f f' s x) : - has_deriv_within_at (λ x, c - f x) (-f') s x := -hf.const_sub c - -theorem has_strict_deriv_at.const_sub (c : F) (hf : has_strict_deriv_at f f' x) : - has_strict_deriv_at (λ x, c - f x) (-f') x := -by simpa only [sub_eq_add_neg] using hf.neg.const_add c - -theorem has_deriv_at.const_sub (c : F) (hf : has_deriv_at f f' x) : - has_deriv_at (λ x, c - f x) (-f') x := -hf.const_sub c - -lemma deriv_within_const_sub (hxs : unique_diff_within_at 𝕜 s x) (c : F) : - deriv_within (λy, c - f y) s x = -deriv_within f s x := -by simp [deriv_within, fderiv_within_const_sub hxs] - -lemma deriv_const_sub (c : F) : deriv (λ y, c - f y) x = -deriv f x := -by simp only [← deriv_within_univ, - deriv_within_const_sub (unique_diff_within_at_univ : unique_diff_within_at 𝕜 _ _)] - -end sub - -section continuous -/-! ### Continuity of a function admitting a derivative -/ - -theorem has_deriv_at_filter.tendsto_nhds - (hL : L ≤ 𝓝 x) (h : has_deriv_at_filter f f' x L) : - tendsto f L (𝓝 (f x)) := -h.tendsto_nhds hL - -theorem has_deriv_within_at.continuous_within_at - (h : has_deriv_within_at f f' s x) : continuous_within_at f s x := -has_deriv_at_filter.tendsto_nhds inf_le_left h - -theorem has_deriv_at.continuous_at (h : has_deriv_at f f' x) : continuous_at f x := -has_deriv_at_filter.tendsto_nhds le_rfl h - -protected theorem has_deriv_at.continuous_on {f f' : 𝕜 → F} - (hderiv : ∀ x ∈ s, has_deriv_at f (f' x) x) : continuous_on f s := -λ x hx, (hderiv x hx).continuous_at.continuous_within_at - -end continuous - -section cartesian_product -/-! ### Derivative of the cartesian product of two functions -/ - -variables {G : Type w} [normed_add_comm_group G] [normed_space 𝕜 G] -variables {f₂ : 𝕜 → G} {f₂' : G} - -lemma has_deriv_at_filter.prod - (hf₁ : has_deriv_at_filter f₁ f₁' x L) (hf₂ : has_deriv_at_filter f₂ f₂' x L) : - has_deriv_at_filter (λ x, (f₁ x, f₂ x)) (f₁', f₂') x L := -hf₁.prod hf₂ - -lemma has_deriv_within_at.prod - (hf₁ : has_deriv_within_at f₁ f₁' s x) (hf₂ : has_deriv_within_at f₂ f₂' s x) : - has_deriv_within_at (λ x, (f₁ x, f₂ x)) (f₁', f₂') s x := -hf₁.prod hf₂ - -lemma has_deriv_at.prod (hf₁ : has_deriv_at f₁ f₁' x) (hf₂ : has_deriv_at f₂ f₂' x) : - has_deriv_at (λ x, (f₁ x, f₂ x)) (f₁', f₂') x := -hf₁.prod hf₂ - -lemma has_strict_deriv_at.prod (hf₁ : has_strict_deriv_at f₁ f₁' x) - (hf₂ : has_strict_deriv_at f₂ f₂' x) : - has_strict_deriv_at (λ x, (f₁ x, f₂ x)) (f₁', f₂') x := -hf₁.prod hf₂ - -end cartesian_product - -section composition -/-! -### Derivative of the composition of a vector function and a scalar function - -We use `scomp` in lemmas on composition of vector valued and scalar valued functions, and `comp` -in lemmas on composition of scalar valued functions, in analogy for `smul` and `mul` (and also -because the `comp` version with the shorter name will show up much more often in applications). -The formula for the derivative involves `smul` in `scomp` lemmas, which can be reduced to -usual multiplication in `comp` lemmas. --/ - -/- For composition lemmas, we put x explicit to help the elaborator, as otherwise Lean tends to -get confused since there are too many possibilities for composition -/ -variables {𝕜' : Type*} [nontrivially_normed_field 𝕜'] [normed_algebra 𝕜 𝕜'] - [normed_space 𝕜' F] [is_scalar_tower 𝕜 𝕜' F] {s' t' : set 𝕜'} - {h : 𝕜 → 𝕜'} {h₁ : 𝕜 → 𝕜} {h₂ : 𝕜' → 𝕜'} {h' h₂' : 𝕜'} {h₁' : 𝕜} - {g₁ : 𝕜' → F} {g₁' : F} {L' : filter 𝕜'} (x) - -theorem has_deriv_at_filter.scomp - (hg : has_deriv_at_filter g₁ g₁' (h x) L') - (hh : has_deriv_at_filter h h' x L) (hL : tendsto h L L'): - has_deriv_at_filter (g₁ ∘ h) (h' • g₁') x L := -by simpa using ((hg.restrict_scalars 𝕜).comp x hh hL).has_deriv_at_filter - -theorem has_deriv_within_at.scomp_has_deriv_at - (hg : has_deriv_within_at g₁ g₁' s' (h x)) - (hh : has_deriv_at h h' x) (hs : ∀ x, h x ∈ s') : - has_deriv_at (g₁ ∘ h) (h' • g₁') x := -hg.scomp x hh $ tendsto_inf.2 ⟨hh.continuous_at, tendsto_principal.2 $ eventually_of_forall hs⟩ - -theorem has_deriv_within_at.scomp - (hg : has_deriv_within_at g₁ g₁' t' (h x)) - (hh : has_deriv_within_at h h' s x) (hst : maps_to h s t') : - has_deriv_within_at (g₁ ∘ h) (h' • g₁') s x := -hg.scomp x hh $ hh.continuous_within_at.tendsto_nhds_within hst - -/-- The chain rule. -/ -theorem has_deriv_at.scomp - (hg : has_deriv_at g₁ g₁' (h x)) (hh : has_deriv_at h h' x) : - has_deriv_at (g₁ ∘ h) (h' • g₁') x := -hg.scomp x hh hh.continuous_at - -theorem has_strict_deriv_at.scomp - (hg : has_strict_deriv_at g₁ g₁' (h x)) (hh : has_strict_deriv_at h h' x) : - has_strict_deriv_at (g₁ ∘ h) (h' • g₁') x := -by simpa using ((hg.restrict_scalars 𝕜).comp x hh).has_strict_deriv_at - -theorem has_deriv_at.scomp_has_deriv_within_at - (hg : has_deriv_at g₁ g₁' (h x)) (hh : has_deriv_within_at h h' s x) : - has_deriv_within_at (g₁ ∘ h) (h' • g₁') s x := -has_deriv_within_at.scomp x hg.has_deriv_within_at hh (maps_to_univ _ _) - -lemma deriv_within.scomp - (hg : differentiable_within_at 𝕜' g₁ t' (h x)) (hh : differentiable_within_at 𝕜 h s x) - (hs : maps_to h s t') (hxs : unique_diff_within_at 𝕜 s x) : - deriv_within (g₁ ∘ h) s x = deriv_within h s x • deriv_within g₁ t' (h x) := -(has_deriv_within_at.scomp x hg.has_deriv_within_at hh.has_deriv_within_at hs).deriv_within hxs - -lemma deriv.scomp - (hg : differentiable_at 𝕜' g₁ (h x)) (hh : differentiable_at 𝕜 h x) : - deriv (g₁ ∘ h) x = deriv h x • deriv g₁ (h x) := -(has_deriv_at.scomp x hg.has_deriv_at hh.has_deriv_at).deriv - -/-! ### Derivative of the composition of a scalar and vector functions -/ - -theorem has_deriv_at_filter.comp_has_fderiv_at_filter {f : E → 𝕜'} {f' : E →L[𝕜] 𝕜'} (x) - {L'' : filter E} (hh₂ : has_deriv_at_filter h₂ h₂' (f x) L') - (hf : has_fderiv_at_filter f f' x L'') (hL : tendsto f L'' L') : - has_fderiv_at_filter (h₂ ∘ f) (h₂' • f') x L'' := -by { convert (hh₂.restrict_scalars 𝕜).comp x hf hL, ext x, simp [mul_comm] } - -theorem has_strict_deriv_at.comp_has_strict_fderiv_at {f : E → 𝕜'} {f' : E →L[𝕜] 𝕜'} (x) - (hh : has_strict_deriv_at h₂ h₂' (f x)) (hf : has_strict_fderiv_at f f' x) : - has_strict_fderiv_at (h₂ ∘ f) (h₂' • f') x := -begin - rw has_strict_deriv_at at hh, - convert (hh.restrict_scalars 𝕜).comp x hf, - ext x, - simp [mul_comm] -end - -theorem has_deriv_at.comp_has_fderiv_at {f : E → 𝕜'} {f' : E →L[𝕜] 𝕜'} (x) - (hh : has_deriv_at h₂ h₂' (f x)) (hf : has_fderiv_at f f' x) : - has_fderiv_at (h₂ ∘ f) (h₂' • f') x := -hh.comp_has_fderiv_at_filter x hf hf.continuous_at - -theorem has_deriv_at.comp_has_fderiv_within_at {f : E → 𝕜'} {f' : E →L[𝕜] 𝕜'} {s} (x) - (hh : has_deriv_at h₂ h₂' (f x)) (hf : has_fderiv_within_at f f' s x) : - has_fderiv_within_at (h₂ ∘ f) (h₂' • f') s x := -hh.comp_has_fderiv_at_filter x hf hf.continuous_within_at - -theorem has_deriv_within_at.comp_has_fderiv_within_at {f : E → 𝕜'} {f' : E →L[𝕜] 𝕜'} {s t} (x) - (hh : has_deriv_within_at h₂ h₂' t (f x)) (hf : has_fderiv_within_at f f' s x) - (hst : maps_to f s t) : - has_fderiv_within_at (h₂ ∘ f) (h₂' • f') s x := -hh.comp_has_fderiv_at_filter x hf $ hf.continuous_within_at.tendsto_nhds_within hst - -/-! ### Derivative of the composition of two scalar functions -/ - -theorem has_deriv_at_filter.comp - (hh₂ : has_deriv_at_filter h₂ h₂' (h x) L') - (hh : has_deriv_at_filter h h' x L) (hL : tendsto h L L') : - has_deriv_at_filter (h₂ ∘ h) (h₂' * h') x L := -by { rw mul_comm, exact hh₂.scomp x hh hL } - -theorem has_deriv_within_at.comp - (hh₂ : has_deriv_within_at h₂ h₂' s' (h x)) - (hh : has_deriv_within_at h h' s x) (hst : maps_to h s s') : - has_deriv_within_at (h₂ ∘ h) (h₂' * h') s x := -by { rw mul_comm, exact hh₂.scomp x hh hst, } - -/-- The chain rule. -/ -theorem has_deriv_at.comp - (hh₂ : has_deriv_at h₂ h₂' (h x)) (hh : has_deriv_at h h' x) : - has_deriv_at (h₂ ∘ h) (h₂' * h') x := -hh₂.comp x hh hh.continuous_at - -theorem has_strict_deriv_at.comp - (hh₂ : has_strict_deriv_at h₂ h₂' (h x)) (hh : has_strict_deriv_at h h' x) : - has_strict_deriv_at (h₂ ∘ h) (h₂' * h') x := -by { rw mul_comm, exact hh₂.scomp x hh } - -theorem has_deriv_at.comp_has_deriv_within_at - (hh₂ : has_deriv_at h₂ h₂' (h x)) (hh : has_deriv_within_at h h' s x) : - has_deriv_within_at (h₂ ∘ h) (h₂' * h') s x := -hh₂.has_deriv_within_at.comp x hh (maps_to_univ _ _) - -lemma deriv_within.comp - (hh₂ : differentiable_within_at 𝕜' h₂ s' (h x)) (hh : differentiable_within_at 𝕜 h s x) - (hs : maps_to h s s') (hxs : unique_diff_within_at 𝕜 s x) : - deriv_within (h₂ ∘ h) s x = deriv_within h₂ s' (h x) * deriv_within h s x := -(hh₂.has_deriv_within_at.comp x hh.has_deriv_within_at hs).deriv_within hxs - -lemma deriv.comp - (hh₂ : differentiable_at 𝕜' h₂ (h x)) (hh : differentiable_at 𝕜 h x) : - deriv (h₂ ∘ h) x = deriv h₂ (h x) * deriv h x := -(hh₂.has_deriv_at.comp x hh.has_deriv_at).deriv - -protected lemma has_deriv_at_filter.iterate {f : 𝕜 → 𝕜} {f' : 𝕜} - (hf : has_deriv_at_filter f f' x L) (hL : tendsto f L L) (hx : f x = x) (n : ℕ) : - has_deriv_at_filter (f^[n]) (f'^n) x L := -begin - have := hf.iterate hL hx n, - rwa [continuous_linear_map.smul_right_one_pow] at this -end - -protected lemma has_deriv_at.iterate {f : 𝕜 → 𝕜} {f' : 𝕜} - (hf : has_deriv_at f f' x) (hx : f x = x) (n : ℕ) : - has_deriv_at (f^[n]) (f'^n) x := -begin - have := has_fderiv_at.iterate hf hx n, - rwa [continuous_linear_map.smul_right_one_pow] at this -end - -protected lemma has_deriv_within_at.iterate {f : 𝕜 → 𝕜} {f' : 𝕜} - (hf : has_deriv_within_at f f' s x) (hx : f x = x) (hs : maps_to f s s) (n : ℕ) : - has_deriv_within_at (f^[n]) (f'^n) s x := -begin - have := has_fderiv_within_at.iterate hf hx hs n, - rwa [continuous_linear_map.smul_right_one_pow] at this -end - -protected lemma has_strict_deriv_at.iterate {f : 𝕜 → 𝕜} {f' : 𝕜} - (hf : has_strict_deriv_at f f' x) (hx : f x = x) (n : ℕ) : - has_strict_deriv_at (f^[n]) (f'^n) x := -begin - have := hf.iterate hx n, - rwa [continuous_linear_map.smul_right_one_pow] at this -end - -end composition - -section composition_vector -/-! ### Derivative of the composition of a function between vector spaces and a function on `𝕜` -/ - -open continuous_linear_map - -variables {l : F → E} {l' : F →L[𝕜] E} -variable (x) - -/-- The composition `l ∘ f` where `l : F → E` and `f : 𝕜 → F`, has a derivative within a set -equal to the Fréchet derivative of `l` applied to the derivative of `f`. -/ -theorem has_fderiv_within_at.comp_has_deriv_within_at {t : set F} - (hl : has_fderiv_within_at l l' t (f x)) (hf : has_deriv_within_at f f' s x) - (hst : maps_to f s t) : - has_deriv_within_at (l ∘ f) (l' f') s x := -by simpa only [one_apply, one_smul, smul_right_apply, coe_comp', (∘)] - using (hl.comp x hf.has_fderiv_within_at hst).has_deriv_within_at - -theorem has_fderiv_at.comp_has_deriv_within_at - (hl : has_fderiv_at l l' (f x)) (hf : has_deriv_within_at f f' s x) : - has_deriv_within_at (l ∘ f) (l' f') s x := -hl.has_fderiv_within_at.comp_has_deriv_within_at x hf (maps_to_univ _ _) - -/-- The composition `l ∘ f` where `l : F → E` and `f : 𝕜 → F`, has a derivative equal to the -Fréchet derivative of `l` applied to the derivative of `f`. -/ -theorem has_fderiv_at.comp_has_deriv_at (hl : has_fderiv_at l l' (f x)) (hf : has_deriv_at f f' x) : - has_deriv_at (l ∘ f) (l' f') x := -has_deriv_within_at_univ.mp $ hl.comp_has_deriv_within_at x hf.has_deriv_within_at - -theorem has_strict_fderiv_at.comp_has_strict_deriv_at - (hl : has_strict_fderiv_at l l' (f x)) (hf : has_strict_deriv_at f f' x) : - has_strict_deriv_at (l ∘ f) (l' f') x := -by simpa only [one_apply, one_smul, smul_right_apply, coe_comp', (∘)] - using (hl.comp x hf.has_strict_fderiv_at).has_strict_deriv_at - -lemma fderiv_within.comp_deriv_within {t : set F} - (hl : differentiable_within_at 𝕜 l t (f x)) (hf : differentiable_within_at 𝕜 f s x) - (hs : maps_to f s t) (hxs : unique_diff_within_at 𝕜 s x) : - deriv_within (l ∘ f) s x = (fderiv_within 𝕜 l t (f x) : F → E) (deriv_within f s x) := -(hl.has_fderiv_within_at.comp_has_deriv_within_at x hf.has_deriv_within_at hs).deriv_within hxs - -lemma fderiv.comp_deriv - (hl : differentiable_at 𝕜 l (f x)) (hf : differentiable_at 𝕜 f x) : - deriv (l ∘ f) x = (fderiv 𝕜 l (f x) : F → E) (deriv f x) := -(hl.has_fderiv_at.comp_has_deriv_at x hf.has_deriv_at).deriv - -end composition_vector - -section mul -/-! ### Derivative of the multiplication of two functions -/ -variables {𝕜' 𝔸 : Type*} [normed_field 𝕜'] [normed_ring 𝔸] [normed_algebra 𝕜 𝕜'] - [normed_algebra 𝕜 𝔸] {c d : 𝕜 → 𝔸} {c' d' : 𝔸} {u v : 𝕜 → 𝕜'} - -theorem has_deriv_within_at.mul - (hc : has_deriv_within_at c c' s x) (hd : has_deriv_within_at d d' s x) : - has_deriv_within_at (λ y, c y * d y) (c' * d x + c x * d') s x := -begin - have := (has_fderiv_within_at.mul' hc hd).has_deriv_within_at, - rwa [continuous_linear_map.add_apply, continuous_linear_map.smul_apply, - continuous_linear_map.smul_right_apply, continuous_linear_map.smul_right_apply, - continuous_linear_map.smul_right_apply, continuous_linear_map.one_apply, - one_smul, one_smul, add_comm] at this, -end - -theorem has_deriv_at.mul (hc : has_deriv_at c c' x) (hd : has_deriv_at d d' x) : - has_deriv_at (λ y, c y * d y) (c' * d x + c x * d') x := -begin - rw [← has_deriv_within_at_univ] at *, - exact hc.mul hd -end - -theorem has_strict_deriv_at.mul - (hc : has_strict_deriv_at c c' x) (hd : has_strict_deriv_at d d' x) : - has_strict_deriv_at (λ y, c y * d y) (c' * d x + c x * d') x := -begin - have := (has_strict_fderiv_at.mul' hc hd).has_strict_deriv_at, - rwa [continuous_linear_map.add_apply, continuous_linear_map.smul_apply, - continuous_linear_map.smul_right_apply, continuous_linear_map.smul_right_apply, - continuous_linear_map.smul_right_apply, continuous_linear_map.one_apply, - one_smul, one_smul, add_comm] at this, -end - -lemma deriv_within_mul (hxs : unique_diff_within_at 𝕜 s x) - (hc : differentiable_within_at 𝕜 c s x) (hd : differentiable_within_at 𝕜 d s x) : - deriv_within (λ y, c y * d y) s x = deriv_within c s x * d x + c x * deriv_within d s x := -(hc.has_deriv_within_at.mul hd.has_deriv_within_at).deriv_within hxs - -@[simp] lemma deriv_mul (hc : differentiable_at 𝕜 c x) (hd : differentiable_at 𝕜 d x) : - deriv (λ y, c y * d y) x = deriv c x * d x + c x * deriv d x := -(hc.has_deriv_at.mul hd.has_deriv_at).deriv - -theorem has_deriv_within_at.mul_const (hc : has_deriv_within_at c c' s x) (d : 𝔸) : - has_deriv_within_at (λ y, c y * d) (c' * d) s x := -begin - convert hc.mul (has_deriv_within_at_const x s d), - rw [mul_zero, add_zero] -end - -theorem has_deriv_at.mul_const (hc : has_deriv_at c c' x) (d : 𝔸) : - has_deriv_at (λ y, c y * d) (c' * d) x := -begin - rw [← has_deriv_within_at_univ] at *, - exact hc.mul_const d -end - -theorem has_deriv_at_mul_const (c : 𝕜) : has_deriv_at (λ x, x * c) c x := -by simpa only [one_mul] using (has_deriv_at_id' x).mul_const c - -theorem has_strict_deriv_at.mul_const (hc : has_strict_deriv_at c c' x) (d : 𝔸) : - has_strict_deriv_at (λ y, c y * d) (c' * d) x := -begin - convert hc.mul (has_strict_deriv_at_const x d), - rw [mul_zero, add_zero] -end - -lemma deriv_within_mul_const (hxs : unique_diff_within_at 𝕜 s x) - (hc : differentiable_within_at 𝕜 c s x) (d : 𝔸) : - deriv_within (λ y, c y * d) s x = deriv_within c s x * d := -(hc.has_deriv_within_at.mul_const d).deriv_within hxs - -lemma deriv_mul_const (hc : differentiable_at 𝕜 c x) (d : 𝔸) : - deriv (λ y, c y * d) x = deriv c x * d := -(hc.has_deriv_at.mul_const d).deriv - -lemma deriv_mul_const_field (v : 𝕜') : - deriv (λ y, u y * v) x = deriv u x * v := -begin - by_cases hu : differentiable_at 𝕜 u x, - { exact deriv_mul_const hu v }, - { rw [deriv_zero_of_not_differentiable_at hu, zero_mul], - rcases eq_or_ne v 0 with rfl|hd, - { simp only [mul_zero, deriv_const] }, - { refine deriv_zero_of_not_differentiable_at (mt (λ H, _) hu), - simpa only [mul_inv_cancel_right₀ hd] using H.mul_const v⁻¹ } } -end - -@[simp] lemma deriv_mul_const_field' (v : 𝕜') : deriv (λ x, u x * v) = λ x, deriv u x * v := -funext $ λ _, deriv_mul_const_field v - -theorem has_deriv_within_at.const_mul (c : 𝔸) (hd : has_deriv_within_at d d' s x) : - has_deriv_within_at (λ y, c * d y) (c * d') s x := -begin - convert (has_deriv_within_at_const x s c).mul hd, - rw [zero_mul, zero_add] -end - -theorem has_deriv_at.const_mul (c : 𝔸) (hd : has_deriv_at d d' x) : - has_deriv_at (λ y, c * d y) (c * d') x := -begin - rw [← has_deriv_within_at_univ] at *, - exact hd.const_mul c -end - -theorem has_strict_deriv_at.const_mul (c : 𝔸) (hd : has_strict_deriv_at d d' x) : - has_strict_deriv_at (λ y, c * d y) (c * d') x := -begin - convert (has_strict_deriv_at_const _ _).mul hd, - rw [zero_mul, zero_add] -end - -lemma deriv_within_const_mul (hxs : unique_diff_within_at 𝕜 s x) - (c : 𝔸) (hd : differentiable_within_at 𝕜 d s x) : - deriv_within (λ y, c * d y) s x = c * deriv_within d s x := -(hd.has_deriv_within_at.const_mul c).deriv_within hxs - -lemma deriv_const_mul (c : 𝔸) (hd : differentiable_at 𝕜 d x) : - deriv (λ y, c * d y) x = c * deriv d x := -(hd.has_deriv_at.const_mul c).deriv - -lemma deriv_const_mul_field (u : 𝕜') : deriv (λ y, u * v y) x = u * deriv v x := -by simp only [mul_comm u, deriv_mul_const_field] - -@[simp] lemma deriv_const_mul_field' (u : 𝕜') : deriv (λ x, u * v x) = λ x, u * deriv v x := -funext (λ x, deriv_const_mul_field u) - -end mul - -section inverse -/-! ### Derivative of `x ↦ x⁻¹` -/ - -theorem has_strict_deriv_at_inv (hx : x ≠ 0) : has_strict_deriv_at has_inv.inv (-(x^2)⁻¹) x := -begin - suffices : (λ p : 𝕜 × 𝕜, (p.1 - p.2) * ((x * x)⁻¹ - (p.1 * p.2)⁻¹)) =o[𝓝 (x, x)] - (λ p, (p.1 - p.2) * 1), - { refine this.congr' _ (eventually_of_forall $ λ _, mul_one _), - refine eventually.mono (is_open.mem_nhds (is_open_ne.prod is_open_ne) ⟨hx, hx⟩) _, - rintro ⟨y, z⟩ ⟨hy, hz⟩, - simp only [mem_set_of_eq] at hy hz, -- hy : y ≠ 0, hz : z ≠ 0 - field_simp [hx, hy, hz], ring, }, - refine (is_O_refl (λ p : 𝕜 × 𝕜, p.1 - p.2) _).mul_is_o ((is_o_one_iff _).2 _), - rw [← sub_self (x * x)⁻¹], - exact tendsto_const_nhds.sub ((continuous_mul.tendsto (x, x)).inv₀ $ mul_ne_zero hx hx) -end - -theorem has_deriv_at_inv (x_ne_zero : x ≠ 0) : - has_deriv_at (λy, y⁻¹) (-(x^2)⁻¹) x := -(has_strict_deriv_at_inv x_ne_zero).has_deriv_at - -theorem has_deriv_within_at_inv (x_ne_zero : x ≠ 0) (s : set 𝕜) : - has_deriv_within_at (λx, x⁻¹) (-(x^2)⁻¹) s x := -(has_deriv_at_inv x_ne_zero).has_deriv_within_at - -lemma differentiable_at_inv : - differentiable_at 𝕜 (λx, x⁻¹) x ↔ x ≠ 0:= -⟨λ H, normed_field.continuous_at_inv.1 H.continuous_at, - λ H, (has_deriv_at_inv H).differentiable_at⟩ - -lemma differentiable_within_at_inv (x_ne_zero : x ≠ 0) : - differentiable_within_at 𝕜 (λx, x⁻¹) s x := -(differentiable_at_inv.2 x_ne_zero).differentiable_within_at - -lemma differentiable_on_inv : differentiable_on 𝕜 (λx:𝕜, x⁻¹) {x | x ≠ 0} := -λx hx, differentiable_within_at_inv hx - -lemma deriv_inv : deriv (λx, x⁻¹) x = -(x^2)⁻¹ := -begin - rcases eq_or_ne x 0 with rfl|hne, - { simp [deriv_zero_of_not_differentiable_at (mt differentiable_at_inv.1 (not_not.2 rfl))] }, - { exact (has_deriv_at_inv hne).deriv } -end - -@[simp] lemma deriv_inv' : deriv (λ x : 𝕜, x⁻¹) = λ x, -(x ^ 2)⁻¹ := funext (λ x, deriv_inv) - -lemma deriv_within_inv (x_ne_zero : x ≠ 0) (hxs : unique_diff_within_at 𝕜 s x) : - deriv_within (λx, x⁻¹) s x = -(x^2)⁻¹ := -begin - rw differentiable_at.deriv_within (differentiable_at_inv.2 x_ne_zero) hxs, - exact deriv_inv -end - -lemma has_fderiv_at_inv (x_ne_zero : x ≠ 0) : - has_fderiv_at (λx, x⁻¹) (smul_right (1 : 𝕜 →L[𝕜] 𝕜) (-(x^2)⁻¹) : 𝕜 →L[𝕜] 𝕜) x := -has_deriv_at_inv x_ne_zero - -lemma has_fderiv_within_at_inv (x_ne_zero : x ≠ 0) : - has_fderiv_within_at (λx, x⁻¹) (smul_right (1 : 𝕜 →L[𝕜] 𝕜) (-(x^2)⁻¹) : 𝕜 →L[𝕜] 𝕜) s x := -(has_fderiv_at_inv x_ne_zero).has_fderiv_within_at - -lemma fderiv_inv : - fderiv 𝕜 (λx, x⁻¹) x = smul_right (1 : 𝕜 →L[𝕜] 𝕜) (-(x^2)⁻¹) := -by rw [← deriv_fderiv, deriv_inv] - -lemma fderiv_within_inv (x_ne_zero : x ≠ 0) (hxs : unique_diff_within_at 𝕜 s x) : - fderiv_within 𝕜 (λx, x⁻¹) s x = smul_right (1 : 𝕜 →L[𝕜] 𝕜) (-(x^2)⁻¹) := -begin - rw differentiable_at.fderiv_within (differentiable_at_inv.2 x_ne_zero) hxs, - exact fderiv_inv -end - -variables {c : 𝕜 → 𝕜} {h : E → 𝕜} {c' : 𝕜} {z : E} {S : set E} - -lemma has_deriv_within_at.inv - (hc : has_deriv_within_at c c' s x) (hx : c x ≠ 0) : - has_deriv_within_at (λ y, (c y)⁻¹) (- c' / (c x)^2) s x := -begin - convert (has_deriv_at_inv hx).comp_has_deriv_within_at x hc, - field_simp -end - -lemma has_deriv_at.inv (hc : has_deriv_at c c' x) (hx : c x ≠ 0) : - has_deriv_at (λ y, (c y)⁻¹) (- c' / (c x)^2) x := -begin - rw ← has_deriv_within_at_univ at *, - exact hc.inv hx -end - -lemma differentiable_within_at.inv (hf : differentiable_within_at 𝕜 h S z) (hz : h z ≠ 0) : - differentiable_within_at 𝕜 (λx, (h x)⁻¹) S z := -(differentiable_at_inv.mpr hz).comp_differentiable_within_at z hf - -@[simp] lemma differentiable_at.inv (hf : differentiable_at 𝕜 h z) (hz : h z ≠ 0) : - differentiable_at 𝕜 (λx, (h x)⁻¹) z := -(differentiable_at_inv.mpr hz).comp z hf - -lemma differentiable_on.inv (hf : differentiable_on 𝕜 h S) (hz : ∀ x ∈ S, h x ≠ 0) : - differentiable_on 𝕜 (λx, (h x)⁻¹) S := -λx h, (hf x h).inv (hz x h) - -@[simp] lemma differentiable.inv (hf : differentiable 𝕜 h) (hz : ∀ x, h x ≠ 0) : - differentiable 𝕜 (λx, (h x)⁻¹) := -λx, (hf x).inv (hz x) - -lemma deriv_within_inv' (hc : differentiable_within_at 𝕜 c s x) (hx : c x ≠ 0) - (hxs : unique_diff_within_at 𝕜 s x) : - deriv_within (λx, (c x)⁻¹) s x = - (deriv_within c s x) / (c x)^2 := -(hc.has_deriv_within_at.inv hx).deriv_within hxs - -@[simp] lemma deriv_inv'' (hc : differentiable_at 𝕜 c x) (hx : c x ≠ 0) : - deriv (λx, (c x)⁻¹) x = - (deriv c x) / (c x)^2 := -(hc.has_deriv_at.inv hx).deriv - -end inverse - -section division -/-! ### Derivative of `x ↦ c x / d x` -/ - -variables {𝕜' : Type*} [nontrivially_normed_field 𝕜'] [normed_algebra 𝕜 𝕜'] - {c d : 𝕜 → 𝕜'} {c' d' : 𝕜'} - -lemma has_deriv_within_at.div - (hc : has_deriv_within_at c c' s x) (hd : has_deriv_within_at d d' s x) (hx : d x ≠ 0) : - has_deriv_within_at (λ y, c y / d y) ((c' * d x - c x * d') / (d x)^2) s x := -begin - convert hc.mul ((has_deriv_at_inv hx).comp_has_deriv_within_at x hd), - { simp only [div_eq_mul_inv] }, - { field_simp, ring } -end - -lemma has_strict_deriv_at.div (hc : has_strict_deriv_at c c' x) (hd : has_strict_deriv_at d d' x) - (hx : d x ≠ 0) : - has_strict_deriv_at (λ y, c y / d y) ((c' * d x - c x * d') / (d x)^2) x := -begin - convert hc.mul ((has_strict_deriv_at_inv hx).comp x hd), - { simp only [div_eq_mul_inv] }, - { field_simp, ring } -end - -lemma has_deriv_at.div (hc : has_deriv_at c c' x) (hd : has_deriv_at d d' x) (hx : d x ≠ 0) : - has_deriv_at (λ y, c y / d y) ((c' * d x - c x * d') / (d x)^2) x := -begin - rw ← has_deriv_within_at_univ at *, - exact hc.div hd hx -end - -lemma differentiable_within_at.div - (hc : differentiable_within_at 𝕜 c s x) (hd : differentiable_within_at 𝕜 d s x) (hx : d x ≠ 0) : - differentiable_within_at 𝕜 (λx, c x / d x) s x := -((hc.has_deriv_within_at).div (hd.has_deriv_within_at) hx).differentiable_within_at - -@[simp] lemma differentiable_at.div - (hc : differentiable_at 𝕜 c x) (hd : differentiable_at 𝕜 d x) (hx : d x ≠ 0) : - differentiable_at 𝕜 (λx, c x / d x) x := -((hc.has_deriv_at).div (hd.has_deriv_at) hx).differentiable_at - -lemma differentiable_on.div - (hc : differentiable_on 𝕜 c s) (hd : differentiable_on 𝕜 d s) (hx : ∀ x ∈ s, d x ≠ 0) : - differentiable_on 𝕜 (λx, c x / d x) s := -λx h, (hc x h).div (hd x h) (hx x h) - -@[simp] lemma differentiable.div - (hc : differentiable 𝕜 c) (hd : differentiable 𝕜 d) (hx : ∀ x, d x ≠ 0) : -differentiable 𝕜 (λx, c x / d x) := -λx, (hc x).div (hd x) (hx x) - -lemma deriv_within_div - (hc : differentiable_within_at 𝕜 c s x) (hd : differentiable_within_at 𝕜 d s x) (hx : d x ≠ 0) - (hxs : unique_diff_within_at 𝕜 s x) : - deriv_within (λx, c x / d x) s x - = ((deriv_within c s x) * d x - c x * (deriv_within d s x)) / (d x)^2 := -((hc.has_deriv_within_at).div (hd.has_deriv_within_at) hx).deriv_within hxs - -@[simp] lemma deriv_div - (hc : differentiable_at 𝕜 c x) (hd : differentiable_at 𝕜 d x) (hx : d x ≠ 0) : - deriv (λx, c x / d x) x = ((deriv c x) * d x - c x * (deriv d x)) / (d x)^2 := -((hc.has_deriv_at).div (hd.has_deriv_at) hx).deriv - -lemma has_deriv_at.div_const (hc : has_deriv_at c c' x) (d : 𝕜') : - has_deriv_at (λ x, c x / d) (c' / d) x := -by simpa only [div_eq_mul_inv] using hc.mul_const d⁻¹ - -lemma has_deriv_within_at.div_const (hc : has_deriv_within_at c c' s x) (d : 𝕜') : - has_deriv_within_at (λ x, c x / d) (c' / d) s x := -by simpa only [div_eq_mul_inv] using hc.mul_const d⁻¹ - -lemma has_strict_deriv_at.div_const (hc : has_strict_deriv_at c c' x) (d : 𝕜') : - has_strict_deriv_at (λ x, c x / d) (c' / d) x := -by simpa only [div_eq_mul_inv] using hc.mul_const d⁻¹ - -lemma differentiable_within_at.div_const (hc : differentiable_within_at 𝕜 c s x) (d : 𝕜') : - differentiable_within_at 𝕜 (λx, c x / d) s x := -(hc.has_deriv_within_at.div_const _).differentiable_within_at - -@[simp] lemma differentiable_at.div_const (hc : differentiable_at 𝕜 c x) (d : 𝕜') : - differentiable_at 𝕜 (λ x, c x / d) x := -(hc.has_deriv_at.div_const _).differentiable_at - -lemma differentiable_on.div_const (hc : differentiable_on 𝕜 c s) (d : 𝕜') : - differentiable_on 𝕜 (λx, c x / d) s := -λ x hx, (hc x hx).div_const d - -@[simp] lemma differentiable.div_const (hc : differentiable 𝕜 c) (d : 𝕜') : - differentiable 𝕜 (λx, c x / d) := -λ x, (hc x).div_const d - -lemma deriv_within_div_const (hc : differentiable_within_at 𝕜 c s x) (d : 𝕜') - (hxs : unique_diff_within_at 𝕜 s x) : - deriv_within (λx, c x / d) s x = (deriv_within c s x) / d := -by simp [div_eq_inv_mul, deriv_within_const_mul, hc, hxs] - -@[simp] lemma deriv_div_const (d : 𝕜') : - deriv (λx, c x / d) x = (deriv c x) / d := -by simp only [div_eq_mul_inv, deriv_mul_const_field] - -end division - -section star -/-! ### Derivative of `x ↦ star x` -/ - -variables [star_ring 𝕜] [has_trivial_star 𝕜] [star_add_monoid F] [has_continuous_star F] -variable [star_module 𝕜 F] - -protected theorem has_deriv_at_filter.star (h : has_deriv_at_filter f f' x L) : - has_deriv_at_filter (λ x, star (f x)) (star f') x L := -by simpa using h.star.has_deriv_at_filter - -protected theorem has_deriv_within_at.star (h : has_deriv_within_at f f' s x) : - has_deriv_within_at (λ x, star (f x)) (star f') s x := -h.star - -protected theorem has_deriv_at.star (h : has_deriv_at f f' x) : - has_deriv_at (λ x, star (f x)) (star f') x := -h.star - -protected theorem has_strict_deriv_at.star (h : has_strict_deriv_at f f' x) : - has_strict_deriv_at (λ x, star (f x)) (star f') x := -by simpa using h.star.has_strict_deriv_at - -protected lemma deriv_within.star (hxs : unique_diff_within_at 𝕜 s x) : - deriv_within (λ y, star (f y)) s x = star (deriv_within f s x) := -fun_like.congr_fun (fderiv_within_star hxs) _ - -protected lemma deriv.star : deriv (λ y, star (f y)) x = star (deriv f x) := -fun_like.congr_fun fderiv_star _ - -@[simp] protected lemma deriv.star' : deriv (λ y, star (f y)) = (λ x, star (deriv f x)) := -funext $ λ x, deriv.star - -end star - -section clm_comp_apply -/-! ### Derivative of the pointwise composition/application of continuous linear maps -/ - -open continuous_linear_map - -variables {G : Type*} [normed_add_comm_group G] [normed_space 𝕜 G] {c : 𝕜 → F →L[𝕜] G} - {c' : F →L[𝕜] G} {d : 𝕜 → E →L[𝕜] F} {d' : E →L[𝕜] F} {u : 𝕜 → F} {u' : F} - -lemma has_strict_deriv_at.clm_comp (hc : has_strict_deriv_at c c' x) - (hd : has_strict_deriv_at d d' x) : - has_strict_deriv_at (λ y, (c y).comp (d y)) (c'.comp (d x) + (c x).comp d') x := -begin - have := (hc.has_strict_fderiv_at.clm_comp hd.has_strict_fderiv_at).has_strict_deriv_at, - rwa [add_apply, comp_apply, comp_apply, smul_right_apply, smul_right_apply, one_apply, one_smul, - one_smul, add_comm] at this, -end - -lemma has_deriv_within_at.clm_comp (hc : has_deriv_within_at c c' s x) - (hd : has_deriv_within_at d d' s x) : - has_deriv_within_at (λ y, (c y).comp (d y)) (c'.comp (d x) + (c x).comp d') s x := -begin - have := (hc.has_fderiv_within_at.clm_comp hd.has_fderiv_within_at).has_deriv_within_at, - rwa [add_apply, comp_apply, comp_apply, smul_right_apply, smul_right_apply, one_apply, one_smul, - one_smul, add_comm] at this, -end - -lemma has_deriv_at.clm_comp (hc : has_deriv_at c c' x) (hd : has_deriv_at d d' x) : - has_deriv_at (λ y, (c y).comp (d y)) - (c'.comp (d x) + (c x).comp d') x := -begin - rw [← has_deriv_within_at_univ] at *, - exact hc.clm_comp hd -end - -lemma deriv_within_clm_comp (hc : differentiable_within_at 𝕜 c s x) - (hd : differentiable_within_at 𝕜 d s x) (hxs : unique_diff_within_at 𝕜 s x): - deriv_within (λ y, (c y).comp (d y)) s x = - ((deriv_within c s x).comp (d x) + (c x).comp (deriv_within d s x)) := -(hc.has_deriv_within_at.clm_comp hd.has_deriv_within_at).deriv_within hxs - -lemma deriv_clm_comp (hc : differentiable_at 𝕜 c x) (hd : differentiable_at 𝕜 d x) : - deriv (λ y, (c y).comp (d y)) x = - ((deriv c x).comp (d x) + (c x).comp (deriv d x)) := -(hc.has_deriv_at.clm_comp hd.has_deriv_at).deriv - -lemma has_strict_deriv_at.clm_apply (hc : has_strict_deriv_at c c' x) - (hu : has_strict_deriv_at u u' x) : - has_strict_deriv_at (λ y, (c y) (u y)) (c' (u x) + c x u') x := -begin - have := (hc.has_strict_fderiv_at.clm_apply hu.has_strict_fderiv_at).has_strict_deriv_at, - rwa [add_apply, comp_apply, flip_apply, smul_right_apply, smul_right_apply, one_apply, one_smul, - one_smul, add_comm] at this, -end - -lemma has_deriv_within_at.clm_apply (hc : has_deriv_within_at c c' s x) - (hu : has_deriv_within_at u u' s x) : - has_deriv_within_at (λ y, (c y) (u y)) (c' (u x) + c x u') s x := -begin - have := (hc.has_fderiv_within_at.clm_apply hu.has_fderiv_within_at).has_deriv_within_at, - rwa [add_apply, comp_apply, flip_apply, smul_right_apply, smul_right_apply, one_apply, one_smul, - one_smul, add_comm] at this, -end - -lemma has_deriv_at.clm_apply (hc : has_deriv_at c c' x) (hu : has_deriv_at u u' x) : - has_deriv_at (λ y, (c y) (u y)) (c' (u x) + c x u') x := -begin - have := (hc.has_fderiv_at.clm_apply hu.has_fderiv_at).has_deriv_at, - rwa [add_apply, comp_apply, flip_apply, smul_right_apply, smul_right_apply, one_apply, one_smul, - one_smul, add_comm] at this, -end - -lemma deriv_within_clm_apply (hxs : unique_diff_within_at 𝕜 s x) - (hc : differentiable_within_at 𝕜 c s x) (hu : differentiable_within_at 𝕜 u s x) : - deriv_within (λ y, (c y) (u y)) s x = (deriv_within c s x (u x) + c x (deriv_within u s x)) := -(hc.has_deriv_within_at.clm_apply hu.has_deriv_within_at).deriv_within hxs - -lemma deriv_clm_apply (hc : differentiable_at 𝕜 c x) (hu : differentiable_at 𝕜 u x) : - deriv (λ y, (c y) (u y)) x = (deriv c x (u x) + c x (deriv u x)) := -(hc.has_deriv_at.clm_apply hu.has_deriv_at).deriv - -end clm_comp_apply - -theorem has_strict_deriv_at.has_strict_fderiv_at_equiv {f : 𝕜 → 𝕜} {f' x : 𝕜} - (hf : has_strict_deriv_at f f' x) (hf' : f' ≠ 0) : - has_strict_fderiv_at f - (continuous_linear_equiv.units_equiv_aut 𝕜 (units.mk0 f' hf') : 𝕜 →L[𝕜] 𝕜) x := -hf - -theorem has_deriv_at.has_fderiv_at_equiv {f : 𝕜 → 𝕜} {f' x : 𝕜} (hf : has_deriv_at f f' x) - (hf' : f' ≠ 0) : - has_fderiv_at f (continuous_linear_equiv.units_equiv_aut 𝕜 (units.mk0 f' hf') : 𝕜 →L[𝕜] 𝕜) x := -hf - -/-- If `f (g y) = y` for `y` in some neighborhood of `a`, `g` is continuous at `a`, and `f` has an -invertible derivative `f'` at `g a` in the strict sense, then `g` has the derivative `f'⁻¹` at `a` -in the strict sense. - -This is one of the easy parts of the inverse function theorem: it assumes that we already have an -inverse function. -/ -theorem has_strict_deriv_at.of_local_left_inverse {f g : 𝕜 → 𝕜} {f' a : 𝕜} - (hg : continuous_at g a) (hf : has_strict_deriv_at f f' (g a)) (hf' : f' ≠ 0) - (hfg : ∀ᶠ y in 𝓝 a, f (g y) = y) : - has_strict_deriv_at g f'⁻¹ a := -(hf.has_strict_fderiv_at_equiv hf').of_local_left_inverse hg hfg - -/-- If `f` is a local homeomorphism defined on a neighbourhood of `f.symm a`, and `f` has a -nonzero derivative `f'` at `f.symm a` in the strict sense, then `f.symm` has the derivative `f'⁻¹` -at `a` in the strict sense. - -This is one of the easy parts of the inverse function theorem: it assumes that we already have -an inverse function. -/ -lemma local_homeomorph.has_strict_deriv_at_symm (f : local_homeomorph 𝕜 𝕜) {a f' : 𝕜} - (ha : a ∈ f.target) (hf' : f' ≠ 0) (htff' : has_strict_deriv_at f f' (f.symm a)) : - has_strict_deriv_at f.symm f'⁻¹ a := -htff'.of_local_left_inverse (f.symm.continuous_at ha) hf' (f.eventually_right_inverse ha) - -/-- If `f (g y) = y` for `y` in some neighborhood of `a`, `g` is continuous at `a`, and `f` has an -invertible derivative `f'` at `g a`, then `g` has the derivative `f'⁻¹` at `a`. - -This is one of the easy parts of the inverse function theorem: it assumes that we already have -an inverse function. -/ -theorem has_deriv_at.of_local_left_inverse {f g : 𝕜 → 𝕜} {f' a : 𝕜} - (hg : continuous_at g a) (hf : has_deriv_at f f' (g a)) (hf' : f' ≠ 0) - (hfg : ∀ᶠ y in 𝓝 a, f (g y) = y) : - has_deriv_at g f'⁻¹ a := -(hf.has_fderiv_at_equiv hf').of_local_left_inverse hg hfg - -/-- If `f` is a local homeomorphism defined on a neighbourhood of `f.symm a`, and `f` has an -nonzero derivative `f'` at `f.symm a`, then `f.symm` has the derivative `f'⁻¹` at `a`. - -This is one of the easy parts of the inverse function theorem: it assumes that we already have -an inverse function. -/ -lemma local_homeomorph.has_deriv_at_symm (f : local_homeomorph 𝕜 𝕜) {a f' : 𝕜} - (ha : a ∈ f.target) (hf' : f' ≠ 0) (htff' : has_deriv_at f f' (f.symm a)) : - has_deriv_at f.symm f'⁻¹ a := -htff'.of_local_left_inverse (f.symm.continuous_at ha) hf' (f.eventually_right_inverse ha) - -lemma has_deriv_at.eventually_ne (h : has_deriv_at f f' x) (hf' : f' ≠ 0) : - ∀ᶠ z in 𝓝[≠] x, f z ≠ f x := -(has_deriv_at_iff_has_fderiv_at.1 h).eventually_ne - ⟨‖f'‖⁻¹, λ z, by field_simp [norm_smul, mt norm_eq_zero.1 hf']⟩ - -lemma has_deriv_at.tendsto_punctured_nhds (h : has_deriv_at f f' x) (hf' : f' ≠ 0) : - tendsto f (𝓝[≠] x) (𝓝[≠] (f x)) := -tendsto_nhds_within_of_tendsto_nhds_of_eventually_within _ - h.continuous_at.continuous_within_at (h.eventually_ne hf') - -theorem not_differentiable_within_at_of_local_left_inverse_has_deriv_within_at_zero - {f g : 𝕜 → 𝕜} {a : 𝕜} {s t : set 𝕜} (ha : a ∈ s) (hsu : unique_diff_within_at 𝕜 s a) - (hf : has_deriv_within_at f 0 t (g a)) (hst : maps_to g s t) (hfg : f ∘ g =ᶠ[𝓝[s] a] id) : - ¬differentiable_within_at 𝕜 g s a := -begin - intro hg, - have := (hf.comp a hg.has_deriv_within_at hst).congr_of_eventually_eq_of_mem hfg.symm ha, - simpa using hsu.eq_deriv _ this (has_deriv_within_at_id _ _) -end - -theorem not_differentiable_at_of_local_left_inverse_has_deriv_at_zero - {f g : 𝕜 → 𝕜} {a : 𝕜} (hf : has_deriv_at f 0 (g a)) (hfg : f ∘ g =ᶠ[𝓝 a] id) : - ¬differentiable_at 𝕜 g a := -begin - intro hg, - have := (hf.comp a hg.has_deriv_at).congr_of_eventually_eq hfg.symm, - simpa using this.unique (has_deriv_at_id a) -end - -end - -namespace polynomial -/-! ### Derivative of a polynomial -/ - -variables {R : Type*} [comm_semiring R] [algebra R 𝕜] -variables (p : 𝕜[X]) (q : R[X]) {x : 𝕜} {s : set 𝕜} - -/-- The derivative (in the analysis sense) of a polynomial `p` is given by `p.derivative`. -/ -protected lemma has_strict_deriv_at (x : 𝕜) : - has_strict_deriv_at (λx, p.eval x) (p.derivative.eval x) x := -begin - apply p.induction_on, - { simp [has_strict_deriv_at_const] }, - { assume p q hp hq, - convert hp.add hq; - simp }, - { assume n a h, - convert h.mul (has_strict_deriv_at_id x), - { ext y, simp [pow_add, mul_assoc] }, - { simp only [pow_add, pow_one, derivative_mul, derivative_C, zero_mul, derivative_X_pow, - derivative_X, mul_one, zero_add, eval_mul, eval_C, eval_add, eval_nat_cast, eval_pow, eval_X, - id.def], ring } } -end - -protected lemma has_strict_deriv_at_aeval (x : 𝕜) : - has_strict_deriv_at (λx, aeval x q) (aeval x q.derivative) x := -by simpa only [aeval_def, eval₂_eq_eval_map, derivative_map] - using (q.map (algebra_map R 𝕜)).has_strict_deriv_at x - -/-- The derivative (in the analysis sense) of a polynomial `p` is given by `p.derivative`. -/ -protected lemma has_deriv_at (x : 𝕜) : has_deriv_at (λx, p.eval x) (p.derivative.eval x) x := -(p.has_strict_deriv_at x).has_deriv_at - -protected lemma has_deriv_at_aeval (x : 𝕜) : - has_deriv_at (λx, aeval x q) (aeval x q.derivative) x := -(q.has_strict_deriv_at_aeval x).has_deriv_at - -protected theorem has_deriv_within_at (x : 𝕜) (s : set 𝕜) : - has_deriv_within_at (λx, p.eval x) (p.derivative.eval x) s x := -(p.has_deriv_at x).has_deriv_within_at - -protected theorem has_deriv_within_at_aeval (x : 𝕜) (s : set 𝕜) : - has_deriv_within_at (λx, aeval x q) (aeval x q.derivative) s x := -(q.has_deriv_at_aeval x).has_deriv_within_at - -protected lemma differentiable_at : differentiable_at 𝕜 (λx, p.eval x) x := -(p.has_deriv_at x).differentiable_at - -protected lemma differentiable_at_aeval : differentiable_at 𝕜 (λx, aeval x q) x := -(q.has_deriv_at_aeval x).differentiable_at - -protected lemma differentiable_within_at : differentiable_within_at 𝕜 (λx, p.eval x) s x := -p.differentiable_at.differentiable_within_at - -protected lemma differentiable_within_at_aeval : differentiable_within_at 𝕜 (λx, aeval x q) s x := -q.differentiable_at_aeval.differentiable_within_at - -protected lemma differentiable : differentiable 𝕜 (λx, p.eval x) := -λx, p.differentiable_at - -protected lemma differentiable_aeval : differentiable 𝕜 (λ x : 𝕜, aeval x q) := -λx, q.differentiable_at_aeval - -protected lemma differentiable_on : differentiable_on 𝕜 (λx, p.eval x) s := -p.differentiable.differentiable_on - -protected lemma differentiable_on_aeval : differentiable_on 𝕜 (λx, aeval x q) s := -q.differentiable_aeval.differentiable_on - -@[simp] protected lemma deriv : deriv (λx, p.eval x) x = p.derivative.eval x := -(p.has_deriv_at x).deriv - -@[simp] protected lemma deriv_aeval : deriv (λx, aeval x q) x = aeval x q.derivative := -(q.has_deriv_at_aeval x).deriv - -protected lemma deriv_within (hxs : unique_diff_within_at 𝕜 s x) : - deriv_within (λx, p.eval x) s x = p.derivative.eval x := -begin - rw differentiable_at.deriv_within p.differentiable_at hxs, - exact p.deriv -end - -protected lemma deriv_within_aeval (hxs : unique_diff_within_at 𝕜 s x) : - deriv_within (λx, aeval x q) s x = aeval x q.derivative := -by simpa only [aeval_def, eval₂_eq_eval_map, derivative_map] - using (q.map (algebra_map R 𝕜)).deriv_within hxs - -protected lemma has_fderiv_at (x : 𝕜) : - has_fderiv_at (λx, p.eval x) (smul_right (1 : 𝕜 →L[𝕜] 𝕜) (p.derivative.eval x)) x := -p.has_deriv_at x - -protected lemma has_fderiv_at_aeval (x : 𝕜) : - has_fderiv_at (λx, aeval x q) (smul_right (1 : 𝕜 →L[𝕜] 𝕜) (aeval x q.derivative)) x := -q.has_deriv_at_aeval x - -protected lemma has_fderiv_within_at (x : 𝕜) : - has_fderiv_within_at (λx, p.eval x) (smul_right (1 : 𝕜 →L[𝕜] 𝕜) (p.derivative.eval x)) s x := -(p.has_fderiv_at x).has_fderiv_within_at - -protected lemma has_fderiv_within_at_aeval (x : 𝕜) : - has_fderiv_within_at (λx, aeval x q) (smul_right (1 : 𝕜 →L[𝕜] 𝕜) (aeval x q.derivative)) s x := -(q.has_fderiv_at_aeval x).has_fderiv_within_at - -@[simp] protected lemma fderiv : - fderiv 𝕜 (λx, p.eval x) x = smul_right (1 : 𝕜 →L[𝕜] 𝕜) (p.derivative.eval x) := -(p.has_fderiv_at x).fderiv - -@[simp] protected lemma fderiv_aeval : - fderiv 𝕜 (λx, aeval x q) x = smul_right (1 : 𝕜 →L[𝕜] 𝕜) (aeval x q.derivative) := -(q.has_fderiv_at_aeval x).fderiv - -protected lemma fderiv_within (hxs : unique_diff_within_at 𝕜 s x) : - fderiv_within 𝕜 (λx, p.eval x) s x = smul_right (1 : 𝕜 →L[𝕜] 𝕜) (p.derivative.eval x) := -(p.has_fderiv_within_at x).fderiv_within hxs - -protected lemma fderiv_within_aeval (hxs : unique_diff_within_at 𝕜 s x) : - fderiv_within 𝕜 (λx, aeval x q) s x = smul_right (1 : 𝕜 →L[𝕜] 𝕜) (aeval x q.derivative) := -(q.has_fderiv_within_at_aeval x).fderiv_within hxs - -end polynomial - -section pow -/-! ### Derivative of `x ↦ x^n` for `n : ℕ` -/ -variables {x : 𝕜} {s : set 𝕜} {c : 𝕜 → 𝕜} {c' : 𝕜} -variable (n : ℕ) - -lemma has_strict_deriv_at_pow (n : ℕ) (x : 𝕜) : - has_strict_deriv_at (λx, x^n) ((n : 𝕜) * x^(n-1)) x := -begin - convert (polynomial.C (1 : 𝕜) * (polynomial.X)^n).has_strict_deriv_at x, - { simp }, - { rw [polynomial.derivative_C_mul_X_pow], simp } -end - -lemma has_deriv_at_pow (n : ℕ) (x : 𝕜) : has_deriv_at (λx, x^n) ((n : 𝕜) * x^(n-1)) x := -(has_strict_deriv_at_pow n x).has_deriv_at - -theorem has_deriv_within_at_pow (n : ℕ) (x : 𝕜) (s : set 𝕜) : - has_deriv_within_at (λx, x^n) ((n : 𝕜) * x^(n-1)) s x := -(has_deriv_at_pow n x).has_deriv_within_at - -lemma differentiable_at_pow : differentiable_at 𝕜 (λx, x^n) x := -(has_deriv_at_pow n x).differentiable_at - -lemma differentiable_within_at_pow : differentiable_within_at 𝕜 (λx, x^n) s x := -(differentiable_at_pow n).differentiable_within_at - -lemma differentiable_pow : differentiable 𝕜 (λx:𝕜, x^n) := -λ x, differentiable_at_pow n - -lemma differentiable_on_pow : differentiable_on 𝕜 (λx, x^n) s := -(differentiable_pow n).differentiable_on - -lemma deriv_pow : deriv (λ x, x^n) x = (n : 𝕜) * x^(n-1) := -(has_deriv_at_pow n x).deriv - -@[simp] lemma deriv_pow' : deriv (λ x, x^n) = λ x, (n : 𝕜) * x^(n-1) := -funext $ λ x, deriv_pow n - -lemma deriv_within_pow (hxs : unique_diff_within_at 𝕜 s x) : - deriv_within (λx, x^n) s x = (n : 𝕜) * x^(n-1) := -(has_deriv_within_at_pow n x s).deriv_within hxs - -lemma has_deriv_within_at.pow (hc : has_deriv_within_at c c' s x) : - has_deriv_within_at (λ y, (c y)^n) ((n : 𝕜) * (c x)^(n-1) * c') s x := -(has_deriv_at_pow n (c x)).comp_has_deriv_within_at x hc - -lemma has_deriv_at.pow (hc : has_deriv_at c c' x) : - has_deriv_at (λ y, (c y)^n) ((n : 𝕜) * (c x)^(n-1) * c') x := -by { rw ← has_deriv_within_at_univ at *, exact hc.pow n } - -lemma deriv_within_pow' (hc : differentiable_within_at 𝕜 c s x) - (hxs : unique_diff_within_at 𝕜 s x) : - deriv_within (λx, (c x)^n) s x = (n : 𝕜) * (c x)^(n-1) * (deriv_within c s x) := -(hc.has_deriv_within_at.pow n).deriv_within hxs - -@[simp] lemma deriv_pow'' (hc : differentiable_at 𝕜 c x) : - deriv (λx, (c x)^n) x = (n : 𝕜) * (c x)^(n-1) * (deriv c x) := -(hc.has_deriv_at.pow n).deriv - -end pow - -section zpow -/-! ### Derivative of `x ↦ x^m` for `m : ℤ` -/ -variables {E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] {x : 𝕜} {s : set 𝕜} {m : ℤ} - -lemma has_strict_deriv_at_zpow (m : ℤ) (x : 𝕜) (h : x ≠ 0 ∨ 0 ≤ m) : - has_strict_deriv_at (λx, x^m) ((m : 𝕜) * x^(m-1)) x := -begin - have : ∀ m : ℤ, 0 < m → has_strict_deriv_at (λx, x^m) ((m:𝕜) * x^(m-1)) x, - { assume m hm, - lift m to ℕ using (le_of_lt hm), - simp only [zpow_coe_nat, int.cast_coe_nat], - convert has_strict_deriv_at_pow _ _ using 2, - rw [← int.coe_nat_one, ← int.coe_nat_sub, zpow_coe_nat], - norm_cast at hm, - exact nat.succ_le_of_lt hm }, - rcases lt_trichotomy m 0 with hm|hm|hm, - { have hx : x ≠ 0, from h.resolve_right hm.not_le, - have := (has_strict_deriv_at_inv _).scomp _ (this (-m) (neg_pos.2 hm)); - [skip, exact zpow_ne_zero_of_ne_zero hx _], - simp only [(∘), zpow_neg, one_div, inv_inv, smul_eq_mul] at this, - convert this using 1, - rw [sq, mul_inv, inv_inv, int.cast_neg, neg_mul, neg_mul_neg, - ← zpow_add₀ hx, mul_assoc, ← zpow_add₀ hx], congr, abel }, - { simp only [hm, zpow_zero, int.cast_zero, zero_mul, has_strict_deriv_at_const] }, - { exact this m hm } -end - -lemma has_deriv_at_zpow (m : ℤ) (x : 𝕜) (h : x ≠ 0 ∨ 0 ≤ m) : - has_deriv_at (λx, x^m) ((m : 𝕜) * x^(m-1)) x := -(has_strict_deriv_at_zpow m x h).has_deriv_at - -theorem has_deriv_within_at_zpow (m : ℤ) (x : 𝕜) (h : x ≠ 0 ∨ 0 ≤ m) (s : set 𝕜) : - has_deriv_within_at (λx, x^m) ((m : 𝕜) * x^(m-1)) s x := -(has_deriv_at_zpow m x h).has_deriv_within_at - -lemma differentiable_at_zpow : differentiable_at 𝕜 (λx, x^m) x ↔ x ≠ 0 ∨ 0 ≤ m := -⟨λ H, normed_field.continuous_at_zpow.1 H.continuous_at, - λ H, (has_deriv_at_zpow m x H).differentiable_at⟩ - -lemma differentiable_within_at_zpow (m : ℤ) (x : 𝕜) (h : x ≠ 0 ∨ 0 ≤ m) : - differentiable_within_at 𝕜 (λx, x^m) s x := -(differentiable_at_zpow.mpr h).differentiable_within_at - -lemma differentiable_on_zpow (m : ℤ) (s : set 𝕜) (h : (0 : 𝕜) ∉ s ∨ 0 ≤ m) : - differentiable_on 𝕜 (λx, x^m) s := -λ x hxs, differentiable_within_at_zpow m x $ h.imp_left $ ne_of_mem_of_not_mem hxs - -lemma deriv_zpow (m : ℤ) (x : 𝕜) : deriv (λ x, x ^ m) x = m * x ^ (m - 1) := -begin - by_cases H : x ≠ 0 ∨ 0 ≤ m, - { exact (has_deriv_at_zpow m x H).deriv }, - { rw deriv_zero_of_not_differentiable_at (mt differentiable_at_zpow.1 H), - push_neg at H, rcases H with ⟨rfl, hm⟩, - rw [zero_zpow _ ((sub_one_lt _).trans hm).ne, mul_zero] } -end - -@[simp] lemma deriv_zpow' (m : ℤ) : deriv (λ x : 𝕜, x ^ m) = λ x, m * x ^ (m - 1) := -funext $ deriv_zpow m - -lemma deriv_within_zpow (hxs : unique_diff_within_at 𝕜 s x) (h : x ≠ 0 ∨ 0 ≤ m) : - deriv_within (λx, x^m) s x = (m : 𝕜) * x^(m-1) := -(has_deriv_within_at_zpow m x h s).deriv_within hxs - -@[simp] lemma iter_deriv_zpow' (m : ℤ) (k : ℕ) : - deriv^[k] (λ x : 𝕜, x ^ m) = λ x, (∏ i in finset.range k, (m - i)) * x ^ (m - k) := -begin - induction k with k ihk, - { simp only [one_mul, int.coe_nat_zero, id, sub_zero, finset.prod_range_zero, - function.iterate_zero] }, - { simp only [function.iterate_succ_apply', ihk, deriv_const_mul_field', deriv_zpow', - finset.prod_range_succ, int.coe_nat_succ, ← sub_sub, int.cast_sub, int.cast_coe_nat, - mul_assoc], } -end - -lemma iter_deriv_zpow (m : ℤ) (x : 𝕜) (k : ℕ) : - deriv^[k] (λ y, y ^ m) x = (∏ i in finset.range k, (m - i)) * x ^ (m - k) := -congr_fun (iter_deriv_zpow' m k) x - -lemma iter_deriv_pow (n : ℕ) (x : 𝕜) (k : ℕ) : - deriv^[k] (λx:𝕜, x^n) x = (∏ i in finset.range k, (n - i)) * x^(n-k) := -begin - simp only [← zpow_coe_nat, iter_deriv_zpow, int.cast_coe_nat], - cases le_or_lt k n with hkn hnk, - { rw int.coe_nat_sub hkn }, - { have : ∏ i in finset.range k, (n - i : 𝕜) = 0, - from finset.prod_eq_zero (finset.mem_range.2 hnk) (sub_self _), - simp only [this, zero_mul] } -end - -@[simp] lemma iter_deriv_pow' (n k : ℕ) : - deriv^[k] (λ x : 𝕜, x ^ n) = λ x, (∏ i in finset.range k, (n - i)) * x ^ (n - k) := -funext $ λ x, iter_deriv_pow n x k - -lemma iter_deriv_inv (k : ℕ) (x : 𝕜) : - deriv^[k] has_inv.inv x = (∏ i in finset.range k, (-1 - i)) * x ^ (-1 - k : ℤ) := -by simpa only [zpow_neg_one, int.cast_neg, int.cast_one] using iter_deriv_zpow (-1) x k - -@[simp] lemma iter_deriv_inv' (k : ℕ) : - deriv^[k] has_inv.inv = λ x : 𝕜, (∏ i in finset.range k, (-1 - i)) * x ^ (-1 - k : ℤ) := -funext (iter_deriv_inv k) - -variables {f : E → 𝕜} {t : set E} {a : E} - -lemma differentiable_within_at.zpow (hf : differentiable_within_at 𝕜 f t a) (h : f a ≠ 0 ∨ 0 ≤ m) : - differentiable_within_at 𝕜 (λ x, f x ^ m) t a := -(differentiable_at_zpow.2 h).comp_differentiable_within_at a hf - -lemma differentiable_at.zpow (hf : differentiable_at 𝕜 f a) (h : f a ≠ 0 ∨ 0 ≤ m) : - differentiable_at 𝕜 (λ x, f x ^ m) a := -(differentiable_at_zpow.2 h).comp a hf - -lemma differentiable_on.zpow (hf : differentiable_on 𝕜 f t) (h : (∀ x ∈ t, f x ≠ 0) ∨ 0 ≤ m) : - differentiable_on 𝕜 (λ x, f x ^ m) t := -λ x hx, (hf x hx).zpow $ h.imp_left (λ h, h x hx) - -lemma differentiable.zpow (hf : differentiable 𝕜 f) (h : (∀ x, f x ≠ 0) ∨ 0 ≤ m) : - differentiable 𝕜 (λ x, f x ^ m) := -λ x, (hf x).zpow $ h.imp_left (λ h, h x) - -end zpow - -/-! ### Support of derivatives -/ - -section support - -open function -variables {F : Type*} [normed_add_comm_group F] [normed_space 𝕜 F] {f : 𝕜 → F} - -lemma support_deriv_subset : support (deriv f) ⊆ tsupport f := -begin - intros x, - rw [← not_imp_not], - intro h2x, - rw [not_mem_tsupport_iff_eventually_eq] at h2x, - exact nmem_support.mpr (h2x.deriv_eq.trans (deriv_const x 0)) -end - -lemma has_compact_support.deriv (hf : has_compact_support f) : has_compact_support (deriv f) := -hf.mono' support_deriv_subset - -end support - -/-! ### Upper estimates on liminf and limsup -/ - -section real - -variables {f : ℝ → ℝ} {f' : ℝ} {s : set ℝ} {x : ℝ} {r : ℝ} - -lemma has_deriv_within_at.limsup_slope_le (hf : has_deriv_within_at f f' s x) (hr : f' < r) : - ∀ᶠ z in 𝓝[s \ {x}] x, slope f x z < r := -has_deriv_within_at_iff_tendsto_slope.1 hf (is_open.mem_nhds is_open_Iio hr) - -lemma has_deriv_within_at.limsup_slope_le' (hf : has_deriv_within_at f f' s x) - (hs : x ∉ s) (hr : f' < r) : - ∀ᶠ z in 𝓝[s] x, slope f x z < r := -(has_deriv_within_at_iff_tendsto_slope' hs).1 hf (is_open.mem_nhds is_open_Iio hr) - -lemma has_deriv_within_at.liminf_right_slope_le - (hf : has_deriv_within_at f f' (Ici x) x) (hr : f' < r) : - ∃ᶠ z in 𝓝[>] x, slope f x z < r := -(hf.Ioi_of_Ici.limsup_slope_le' (lt_irrefl x) hr).frequently - -end real - -section real_space - -open metric - -variables {E : Type u} [normed_add_comm_group E] [normed_space ℝ E] {f : ℝ → E} {f' : E} {s : set ℝ} - {x r : ℝ} - -/-- If `f` has derivative `f'` within `s` at `x`, then for any `r > ‖f'‖` the ratio -`‖f z - f x‖ / ‖z - x‖` is less than `r` in some neighborhood of `x` within `s`. -In other words, the limit superior of this ratio as `z` tends to `x` along `s` -is less than or equal to `‖f'‖`. -/ -lemma has_deriv_within_at.limsup_norm_slope_le - (hf : has_deriv_within_at f f' s x) (hr : ‖f'‖ < r) : - ∀ᶠ z in 𝓝[s] x, ‖z - x‖⁻¹ * ‖f z - f x‖ < r := -begin - have hr₀ : 0 < r, from lt_of_le_of_lt (norm_nonneg f') hr, - have A : ∀ᶠ z in 𝓝[s \ {x}] x, ‖(z - x)⁻¹ • (f z - f x)‖ ∈ Iio r, - from (has_deriv_within_at_iff_tendsto_slope.1 hf).norm (is_open.mem_nhds is_open_Iio hr), - have B : ∀ᶠ z in 𝓝[{x}] x, ‖(z - x)⁻¹ • (f z - f x)‖ ∈ Iio r, - from mem_of_superset self_mem_nhds_within - (singleton_subset_iff.2 $ by simp [hr₀]), - have C := mem_sup.2 ⟨A, B⟩, - rw [← nhds_within_union, diff_union_self, nhds_within_union, mem_sup] at C, - filter_upwards [C.1], - simp only [norm_smul, mem_Iio, norm_inv], - exact λ _, id -end - -/-- If `f` has derivative `f'` within `s` at `x`, then for any `r > ‖f'‖` the ratio -`(‖f z‖ - ‖f x‖) / ‖z - x‖` is less than `r` in some neighborhood of `x` within `s`. -In other words, the limit superior of this ratio as `z` tends to `x` along `s` -is less than or equal to `‖f'‖`. - -This lemma is a weaker version of `has_deriv_within_at.limsup_norm_slope_le` -where `‖f z‖ - ‖f x‖` is replaced by `‖f z - f x‖`. -/ -lemma has_deriv_within_at.limsup_slope_norm_le - (hf : has_deriv_within_at f f' s x) (hr : ‖f'‖ < r) : - ∀ᶠ z in 𝓝[s] x, ‖z - x‖⁻¹ * (‖f z‖ - ‖f x‖) < r := -begin - apply (hf.limsup_norm_slope_le hr).mono, - assume z hz, - refine lt_of_le_of_lt (mul_le_mul_of_nonneg_left (norm_sub_norm_le _ _) _) hz, - exact inv_nonneg.2 (norm_nonneg _) -end - -/-- If `f` has derivative `f'` within `(x, +∞)` at `x`, then for any `r > ‖f'‖` the ratio -`‖f z - f x‖ / ‖z - x‖` is frequently less than `r` as `z → x+0`. -In other words, the limit inferior of this ratio as `z` tends to `x+0` -is less than or equal to `‖f'‖`. See also `has_deriv_within_at.limsup_norm_slope_le` -for a stronger version using limit superior and any set `s`. -/ -lemma has_deriv_within_at.liminf_right_norm_slope_le - (hf : has_deriv_within_at f f' (Ici x) x) (hr : ‖f'‖ < r) : - ∃ᶠ z in 𝓝[>] x, ‖z - x‖⁻¹ * ‖f z - f x‖ < r := -(hf.Ioi_of_Ici.limsup_norm_slope_le hr).frequently - -/-- If `f` has derivative `f'` within `(x, +∞)` at `x`, then for any `r > ‖f'‖` the ratio -`(‖f z‖ - ‖f x‖) / (z - x)` is frequently less than `r` as `z → x+0`. -In other words, the limit inferior of this ratio as `z` tends to `x+0` -is less than or equal to `‖f'‖`. - -See also - -* `has_deriv_within_at.limsup_norm_slope_le` for a stronger version using - limit superior and any set `s`; -* `has_deriv_within_at.liminf_right_norm_slope_le` for a stronger version using - `‖f z - f x‖` instead of `‖f z‖ - ‖f x‖`. -/ -lemma has_deriv_within_at.liminf_right_slope_norm_le - (hf : has_deriv_within_at f f' (Ici x) x) (hr : ‖f'‖ < r) : - ∃ᶠ z in 𝓝[>] x, (z - x)⁻¹ * (‖f z‖ - ‖f x‖) < r := -begin - have := (hf.Ioi_of_Ici.limsup_slope_norm_le hr).frequently, - refine this.mp (eventually.mono self_mem_nhds_within _), - assume z hxz hz, - rwa [real.norm_eq_abs, abs_of_pos (sub_pos_of_lt hxz)] at hz -end - -end real_space diff --git a/src/analysis/calculus/deriv/add.lean b/src/analysis/calculus/deriv/add.lean new file mode 100644 index 0000000000000..70f4e35833266 --- /dev/null +++ b/src/analysis/calculus/deriv/add.lean @@ -0,0 +1,304 @@ +/- +Copyright (c) 2019 Gabriel Ebner All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Gabriel Ebner, Sébastien Gouëzel, Yury Kudryashov, Anatole Dedecker +-/ +import analysis.calculus.deriv.basic +import analysis.calculus.fderiv.add + +/-! +# One-dimensional derivatives of sums etc + +In this file we prove formulas about derivatives of `f + g`, `-f`, `f - g`, and `∑ i, f i x` for +functions from the base field to a normed space over this field. + +For a more detailed overview of one-dimensional derivatives in mathlib, see the module docstring of +`analysis/calculus/deriv/basic`. + +## Keywords + +derivative +-/ + +universes u v w +open_locale classical topology big_operators filter ennreal +open filter asymptotics set + +variables {𝕜 : Type u} [nontrivially_normed_field 𝕜] +variables {F : Type v} [normed_add_comm_group F] [normed_space 𝕜 F] +variables {E : Type w} [normed_add_comm_group E] [normed_space 𝕜 E] + +variables {f f₀ f₁ g : 𝕜 → F} +variables {f' f₀' f₁' g' : F} +variables {x : 𝕜} +variables {s t : set 𝕜} +variables {L : filter 𝕜} + +section add +/-! ### Derivative of the sum of two functions -/ + +theorem has_deriv_at_filter.add + (hf : has_deriv_at_filter f f' x L) (hg : has_deriv_at_filter g g' x L) : + has_deriv_at_filter (λ y, f y + g y) (f' + g') x L := +by simpa using (hf.add hg).has_deriv_at_filter + +theorem has_strict_deriv_at.add + (hf : has_strict_deriv_at f f' x) (hg : has_strict_deriv_at g g' x) : + has_strict_deriv_at (λ y, f y + g y) (f' + g') x := +by simpa using (hf.add hg).has_strict_deriv_at + +theorem has_deriv_within_at.add + (hf : has_deriv_within_at f f' s x) (hg : has_deriv_within_at g g' s x) : + has_deriv_within_at (λ y, f y + g y) (f' + g') s x := +hf.add hg + +theorem has_deriv_at.add + (hf : has_deriv_at f f' x) (hg : has_deriv_at g g' x) : + has_deriv_at (λ x, f x + g x) (f' + g') x := +hf.add hg + +lemma deriv_within_add (hxs : unique_diff_within_at 𝕜 s x) + (hf : differentiable_within_at 𝕜 f s x) (hg : differentiable_within_at 𝕜 g s x) : + deriv_within (λy, f y + g y) s x = deriv_within f s x + deriv_within g s x := +(hf.has_deriv_within_at.add hg.has_deriv_within_at).deriv_within hxs + +@[simp] lemma deriv_add + (hf : differentiable_at 𝕜 f x) (hg : differentiable_at 𝕜 g x) : + deriv (λy, f y + g y) x = deriv f x + deriv g x := +(hf.has_deriv_at.add hg.has_deriv_at).deriv + +theorem has_deriv_at_filter.add_const + (hf : has_deriv_at_filter f f' x L) (c : F) : + has_deriv_at_filter (λ y, f y + c) f' x L := +add_zero f' ▸ hf.add (has_deriv_at_filter_const x L c) + +theorem has_deriv_within_at.add_const + (hf : has_deriv_within_at f f' s x) (c : F) : + has_deriv_within_at (λ y, f y + c) f' s x := +hf.add_const c + +theorem has_deriv_at.add_const + (hf : has_deriv_at f f' x) (c : F) : + has_deriv_at (λ x, f x + c) f' x := +hf.add_const c + +lemma deriv_within_add_const (hxs : unique_diff_within_at 𝕜 s x) (c : F) : + deriv_within (λy, f y + c) s x = deriv_within f s x := +by simp only [deriv_within, fderiv_within_add_const hxs] + +lemma deriv_add_const (c : F) : deriv (λy, f y + c) x = deriv f x := +by simp only [deriv, fderiv_add_const] + +@[simp] lemma deriv_add_const' (c : F) : deriv (λ y, f y + c) = deriv f := +funext $ λ x, deriv_add_const c + +theorem has_deriv_at_filter.const_add (c : F) (hf : has_deriv_at_filter f f' x L) : + has_deriv_at_filter (λ y, c + f y) f' x L := +zero_add f' ▸ (has_deriv_at_filter_const x L c).add hf + +theorem has_deriv_within_at.const_add (c : F) (hf : has_deriv_within_at f f' s x) : + has_deriv_within_at (λ y, c + f y) f' s x := +hf.const_add c + +theorem has_deriv_at.const_add (c : F) (hf : has_deriv_at f f' x) : + has_deriv_at (λ x, c + f x) f' x := +hf.const_add c + +lemma deriv_within_const_add (hxs : unique_diff_within_at 𝕜 s x) (c : F) : + deriv_within (λy, c + f y) s x = deriv_within f s x := +by simp only [deriv_within, fderiv_within_const_add hxs] + +lemma deriv_const_add (c : F) : deriv (λy, c + f y) x = deriv f x := +by simp only [deriv, fderiv_const_add] + +@[simp] lemma deriv_const_add' (c : F) : deriv (λ y, c + f y) = deriv f := +funext $ λ x, deriv_const_add c + +end add + +section sum +/-! ### Derivative of a finite sum of functions -/ + +open_locale big_operators + +variables {ι : Type*} {u : finset ι} {A : ι → (𝕜 → F)} {A' : ι → F} + +theorem has_deriv_at_filter.sum (h : ∀ i ∈ u, has_deriv_at_filter (A i) (A' i) x L) : + has_deriv_at_filter (λ y, ∑ i in u, A i y) (∑ i in u, A' i) x L := +by simpa [continuous_linear_map.sum_apply] using (has_fderiv_at_filter.sum h).has_deriv_at_filter + +theorem has_strict_deriv_at.sum (h : ∀ i ∈ u, has_strict_deriv_at (A i) (A' i) x) : + has_strict_deriv_at (λ y, ∑ i in u, A i y) (∑ i in u, A' i) x := +by simpa [continuous_linear_map.sum_apply] using (has_strict_fderiv_at.sum h).has_strict_deriv_at + +theorem has_deriv_within_at.sum (h : ∀ i ∈ u, has_deriv_within_at (A i) (A' i) s x) : + has_deriv_within_at (λ y, ∑ i in u, A i y) (∑ i in u, A' i) s x := +has_deriv_at_filter.sum h + +theorem has_deriv_at.sum (h : ∀ i ∈ u, has_deriv_at (A i) (A' i) x) : + has_deriv_at (λ y, ∑ i in u, A i y) (∑ i in u, A' i) x := +has_deriv_at_filter.sum h + +lemma deriv_within_sum (hxs : unique_diff_within_at 𝕜 s x) + (h : ∀ i ∈ u, differentiable_within_at 𝕜 (A i) s x) : + deriv_within (λ y, ∑ i in u, A i y) s x = ∑ i in u, deriv_within (A i) s x := +(has_deriv_within_at.sum (λ i hi, (h i hi).has_deriv_within_at)).deriv_within hxs + +@[simp] lemma deriv_sum (h : ∀ i ∈ u, differentiable_at 𝕜 (A i) x) : + deriv (λ y, ∑ i in u, A i y) x = ∑ i in u, deriv (A i) x := +(has_deriv_at.sum (λ i hi, (h i hi).has_deriv_at)).deriv + +end sum + + +section neg +/-! ### Derivative of the negative of a function -/ + +theorem has_deriv_at_filter.neg (h : has_deriv_at_filter f f' x L) : + has_deriv_at_filter (λ x, -f x) (-f') x L := +by simpa using h.neg.has_deriv_at_filter + +theorem has_deriv_within_at.neg (h : has_deriv_within_at f f' s x) : + has_deriv_within_at (λ x, -f x) (-f') s x := +h.neg + +theorem has_deriv_at.neg (h : has_deriv_at f f' x) : has_deriv_at (λ x, -f x) (-f') x := +h.neg + +theorem has_strict_deriv_at.neg (h : has_strict_deriv_at f f' x) : + has_strict_deriv_at (λ x, -f x) (-f') x := +by simpa using h.neg.has_strict_deriv_at + +lemma deriv_within.neg (hxs : unique_diff_within_at 𝕜 s x) : + deriv_within (λy, -f y) s x = - deriv_within f s x := +by simp only [deriv_within, fderiv_within_neg hxs, continuous_linear_map.neg_apply] + +lemma deriv.neg : deriv (λy, -f y) x = - deriv f x := +by simp only [deriv, fderiv_neg, continuous_linear_map.neg_apply] + +@[simp] lemma deriv.neg' : deriv (λy, -f y) = (λ x, - deriv f x) := +funext $ λ x, deriv.neg + +end neg + +section neg2 +/-! ### Derivative of the negation function (i.e `has_neg.neg`) -/ + +variables (s x L) + +theorem has_deriv_at_filter_neg : has_deriv_at_filter has_neg.neg (-1) x L := +has_deriv_at_filter.neg $ has_deriv_at_filter_id _ _ + +theorem has_deriv_within_at_neg : has_deriv_within_at has_neg.neg (-1) s x := +has_deriv_at_filter_neg _ _ + +theorem has_deriv_at_neg : has_deriv_at has_neg.neg (-1) x := +has_deriv_at_filter_neg _ _ + +theorem has_deriv_at_neg' : has_deriv_at (λ x, -x) (-1) x := +has_deriv_at_filter_neg _ _ + +theorem has_strict_deriv_at_neg : has_strict_deriv_at has_neg.neg (-1) x := +has_strict_deriv_at.neg $ has_strict_deriv_at_id _ + +lemma deriv_neg : deriv has_neg.neg x = -1 := +has_deriv_at.deriv (has_deriv_at_neg x) + +@[simp] lemma deriv_neg' : deriv (has_neg.neg : 𝕜 → 𝕜) = λ _, -1 := +funext deriv_neg + +@[simp] lemma deriv_neg'' : deriv (λ x : 𝕜, -x) x = -1 := +deriv_neg x + +lemma deriv_within_neg (hxs : unique_diff_within_at 𝕜 s x) : deriv_within has_neg.neg s x = -1 := +(has_deriv_within_at_neg x s).deriv_within hxs + +lemma differentiable_neg : differentiable 𝕜 (has_neg.neg : 𝕜 → 𝕜) := +differentiable.neg differentiable_id + +lemma differentiable_on_neg : differentiable_on 𝕜 (has_neg.neg : 𝕜 → 𝕜) s := +differentiable_on.neg differentiable_on_id + +end neg2 + +section sub +/-! ### Derivative of the difference of two functions -/ + +theorem has_deriv_at_filter.sub + (hf : has_deriv_at_filter f f' x L) (hg : has_deriv_at_filter g g' x L) : + has_deriv_at_filter (λ x, f x - g x) (f' - g') x L := +by simpa only [sub_eq_add_neg] using hf.add hg.neg + +theorem has_deriv_within_at.sub + (hf : has_deriv_within_at f f' s x) (hg : has_deriv_within_at g g' s x) : + has_deriv_within_at (λ x, f x - g x) (f' - g') s x := +hf.sub hg + +theorem has_deriv_at.sub + (hf : has_deriv_at f f' x) (hg : has_deriv_at g g' x) : + has_deriv_at (λ x, f x - g x) (f' - g') x := +hf.sub hg + +theorem has_strict_deriv_at.sub + (hf : has_strict_deriv_at f f' x) (hg : has_strict_deriv_at g g' x) : + has_strict_deriv_at (λ x, f x - g x) (f' - g') x := +by simpa only [sub_eq_add_neg] using hf.add hg.neg + +lemma deriv_within_sub (hxs : unique_diff_within_at 𝕜 s x) + (hf : differentiable_within_at 𝕜 f s x) (hg : differentiable_within_at 𝕜 g s x) : + deriv_within (λy, f y - g y) s x = deriv_within f s x - deriv_within g s x := +(hf.has_deriv_within_at.sub hg.has_deriv_within_at).deriv_within hxs + +@[simp] lemma deriv_sub + (hf : differentiable_at 𝕜 f x) (hg : differentiable_at 𝕜 g x) : + deriv (λ y, f y - g y) x = deriv f x - deriv g x := +(hf.has_deriv_at.sub hg.has_deriv_at).deriv + +theorem has_deriv_at_filter.sub_const + (hf : has_deriv_at_filter f f' x L) (c : F) : + has_deriv_at_filter (λ x, f x - c) f' x L := +by simpa only [sub_eq_add_neg] using hf.add_const (-c) + +theorem has_deriv_within_at.sub_const + (hf : has_deriv_within_at f f' s x) (c : F) : + has_deriv_within_at (λ x, f x - c) f' s x := +hf.sub_const c + +theorem has_deriv_at.sub_const + (hf : has_deriv_at f f' x) (c : F) : + has_deriv_at (λ x, f x - c) f' x := +hf.sub_const c + +lemma deriv_within_sub_const (hxs : unique_diff_within_at 𝕜 s x) (c : F) : + deriv_within (λy, f y - c) s x = deriv_within f s x := +by simp only [deriv_within, fderiv_within_sub_const hxs] + +lemma deriv_sub_const (c : F) : deriv (λ y, f y - c) x = deriv f x := +by simp only [deriv, fderiv_sub_const] + +theorem has_deriv_at_filter.const_sub (c : F) (hf : has_deriv_at_filter f f' x L) : + has_deriv_at_filter (λ x, c - f x) (-f') x L := +by simpa only [sub_eq_add_neg] using hf.neg.const_add c + +theorem has_deriv_within_at.const_sub (c : F) (hf : has_deriv_within_at f f' s x) : + has_deriv_within_at (λ x, c - f x) (-f') s x := +hf.const_sub c + +theorem has_strict_deriv_at.const_sub (c : F) (hf : has_strict_deriv_at f f' x) : + has_strict_deriv_at (λ x, c - f x) (-f') x := +by simpa only [sub_eq_add_neg] using hf.neg.const_add c + +theorem has_deriv_at.const_sub (c : F) (hf : has_deriv_at f f' x) : + has_deriv_at (λ x, c - f x) (-f') x := +hf.const_sub c + +lemma deriv_within_const_sub (hxs : unique_diff_within_at 𝕜 s x) (c : F) : + deriv_within (λy, c - f y) s x = -deriv_within f s x := +by simp [deriv_within, fderiv_within_const_sub hxs] + +lemma deriv_const_sub (c : F) : deriv (λ y, c - f y) x = -deriv f x := +by simp only [← deriv_within_univ, + deriv_within_const_sub (unique_diff_within_at_univ : unique_diff_within_at 𝕜 _ _)] + +end sub + diff --git a/src/analysis/calculus/deriv/basic.lean b/src/analysis/calculus/deriv/basic.lean new file mode 100644 index 0000000000000..cf6c1d7c00b02 --- /dev/null +++ b/src/analysis/calculus/deriv/basic.lean @@ -0,0 +1,599 @@ +/- +Copyright (c) 2019 Gabriel Ebner. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Gabriel Ebner, Sébastien Gouëzel +-/ +import analysis.calculus.fderiv.basic + +/-! + +# One-dimensional derivatives + +This file defines the derivative of a function `f : 𝕜 → F` where `𝕜` is a +normed field and `F` is a normed space over this field. The derivative of +such a function `f` at a point `x` is given by an element `f' : F`. + +The theory is developed analogously to the [Fréchet +derivatives](./fderiv.html). We first introduce predicates defined in terms +of the corresponding predicates for Fréchet derivatives: + + - `has_deriv_at_filter f f' x L` states that the function `f` has the + derivative `f'` at the point `x` as `x` goes along the filter `L`. + + - `has_deriv_within_at f f' s x` states that the function `f` has the + derivative `f'` at the point `x` within the subset `s`. + + - `has_deriv_at f f' x` states that the function `f` has the derivative `f'` + at the point `x`. + + - `has_strict_deriv_at f f' x` states that the function `f` has the derivative `f'` + at the point `x` in the sense of strict differentiability, i.e., + `f y - f z = (y - z) • f' + o (y - z)` as `y, z → x`. + +For the last two notions we also define a functional version: + + - `deriv_within f s x` is a derivative of `f` at `x` within `s`. If the + derivative does not exist, then `deriv_within f s x` equals zero. + + - `deriv f x` is a derivative of `f` at `x`. If the derivative does not + exist, then `deriv f x` equals zero. + +The theorems `fderiv_within_deriv_within` and `fderiv_deriv` show that the +one-dimensional derivatives coincide with the general Fréchet derivatives. + +We also show the existence and compute the derivatives of: + - constants + - the identity function + - linear maps (in `linear.lean`) + - addition (in `add.lean`) + - sum of finitely many functions (in `add.lean`) + - negation (in `add.lean`) + - subtraction (in `add.lean`) + - star (in `star.lean`) + - multiplication of two functions in `𝕜 → 𝕜` (in `mul.lean`) + - multiplication of a function in `𝕜 → 𝕜` and of a function in `𝕜 → E` (in `mul.lean`) + - powers of a function (in `pow.lean` and `zpow.lean`) + - inverse `x → x⁻¹` (in `inv.lean`) + - division (in `inv.lean`) + - composition of a function in `𝕜 → F` with a function in `𝕜 → 𝕜` (in `comp.lean`) + - composition of a function in `F → E` with a function in `𝕜 → F` (in `comp.lean`) + - inverse function (assuming that it exists; the inverse function theorem is in `inverse.lean`) + - polynomials (in `polynomial.lean`) + +For most binary operations we also define `const_op` and `op_const` theorems for the cases when +the first or second argument is a constant. This makes writing chains of `has_deriv_at`'s easier, +and they more frequently lead to the desired result. + +We set up the simplifier so that it can compute the derivative of simple functions. For instance, +```lean +example (x : ℝ) : deriv (λ x, cos (sin x) * exp x) x = (cos(sin(x))-sin(sin(x))*cos(x))*exp(x) := +by { simp, ring } +``` + +## Implementation notes + +Most of the theorems are direct restatements of the corresponding theorems +for Fréchet derivatives. + +The strategy to construct simp lemmas that give the simplifier the possibility to compute +derivatives is the same as the one for differentiability statements, as explained in `fderiv.lean`. +See the explanations there. +-/ + +universes u v w +noncomputable theory +open_locale classical topology big_operators filter ennreal +open filter asymptotics set +open continuous_linear_map (smul_right smul_right_one_eq_iff) + + +variables {𝕜 : Type u} [nontrivially_normed_field 𝕜] +variables {F : Type v} [normed_add_comm_group F] [normed_space 𝕜 F] +variables {E : Type w} [normed_add_comm_group E] [normed_space 𝕜 E] + +/-- +`f` has the derivative `f'` at the point `x` as `x` goes along the filter `L`. + +That is, `f x' = f x + (x' - x) • f' + o(x' - x)` where `x'` converges along the filter `L`. +-/ +def has_deriv_at_filter (f : 𝕜 → F) (f' : F) (x : 𝕜) (L : filter 𝕜) := +has_fderiv_at_filter f (smul_right (1 : 𝕜 →L[𝕜] 𝕜) f') x L + +/-- +`f` has the derivative `f'` at the point `x` within the subset `s`. + +That is, `f x' = f x + (x' - x) • f' + o(x' - x)` where `x'` converges to `x` inside `s`. +-/ +def has_deriv_within_at (f : 𝕜 → F) (f' : F) (s : set 𝕜) (x : 𝕜) := +has_deriv_at_filter f f' x (𝓝[s] x) + +/-- +`f` has the derivative `f'` at the point `x`. + +That is, `f x' = f x + (x' - x) • f' + o(x' - x)` where `x'` converges to `x`. +-/ +def has_deriv_at (f : 𝕜 → F) (f' : F) (x : 𝕜) := +has_deriv_at_filter f f' x (𝓝 x) + +/-- `f` has the derivative `f'` at the point `x` in the sense of strict differentiability. + +That is, `f y - f z = (y - z) • f' + o(y - z)` as `y, z → x`. -/ +def has_strict_deriv_at (f : 𝕜 → F) (f' : F) (x : 𝕜) := +has_strict_fderiv_at f (smul_right (1 : 𝕜 →L[𝕜] 𝕜) f') x + +/-- +Derivative of `f` at the point `x` within the set `s`, if it exists. Zero otherwise. + +If the derivative exists (i.e., `∃ f', has_deriv_within_at f f' s x`), then +`f x' = f x + (x' - x) • deriv_within f s x + o(x' - x)` where `x'` converges to `x` inside `s`. +-/ +def deriv_within (f : 𝕜 → F) (s : set 𝕜) (x : 𝕜) := +fderiv_within 𝕜 f s x 1 + +/-- +Derivative of `f` at the point `x`, if it exists. Zero otherwise. + +If the derivative exists (i.e., `∃ f', has_deriv_at f f' x`), then +`f x' = f x + (x' - x) • deriv f x + o(x' - x)` where `x'` converges to `x`. +-/ +def deriv (f : 𝕜 → F) (x : 𝕜) := +fderiv 𝕜 f x 1 + +variables {f f₀ f₁ g : 𝕜 → F} +variables {f' f₀' f₁' g' : F} +variables {x : 𝕜} +variables {s t : set 𝕜} +variables {L L₁ L₂ : filter 𝕜} + +/-- Expressing `has_fderiv_at_filter f f' x L` in terms of `has_deriv_at_filter` -/ +lemma has_fderiv_at_filter_iff_has_deriv_at_filter {f' : 𝕜 →L[𝕜] F} : + has_fderiv_at_filter f f' x L ↔ has_deriv_at_filter f (f' 1) x L := +by simp [has_deriv_at_filter] + +lemma has_fderiv_at_filter.has_deriv_at_filter {f' : 𝕜 →L[𝕜] F} : + has_fderiv_at_filter f f' x L → has_deriv_at_filter f (f' 1) x L := +has_fderiv_at_filter_iff_has_deriv_at_filter.mp + +/-- Expressing `has_fderiv_within_at f f' s x` in terms of `has_deriv_within_at` -/ +lemma has_fderiv_within_at_iff_has_deriv_within_at {f' : 𝕜 →L[𝕜] F} : + has_fderiv_within_at f f' s x ↔ has_deriv_within_at f (f' 1) s x := +has_fderiv_at_filter_iff_has_deriv_at_filter + +/-- Expressing `has_deriv_within_at f f' s x` in terms of `has_fderiv_within_at` -/ +lemma has_deriv_within_at_iff_has_fderiv_within_at {f' : F} : + has_deriv_within_at f f' s x ↔ + has_fderiv_within_at f (smul_right (1 : 𝕜 →L[𝕜] 𝕜) f') s x := +iff.rfl + +lemma has_fderiv_within_at.has_deriv_within_at {f' : 𝕜 →L[𝕜] F} : + has_fderiv_within_at f f' s x → has_deriv_within_at f (f' 1) s x := +has_fderiv_within_at_iff_has_deriv_within_at.mp + +lemma has_deriv_within_at.has_fderiv_within_at {f' : F} : + has_deriv_within_at f f' s x → has_fderiv_within_at f (smul_right (1 : 𝕜 →L[𝕜] 𝕜) f') s x := +has_deriv_within_at_iff_has_fderiv_within_at.mp + +/-- Expressing `has_fderiv_at f f' x` in terms of `has_deriv_at` -/ +lemma has_fderiv_at_iff_has_deriv_at {f' : 𝕜 →L[𝕜] F} : + has_fderiv_at f f' x ↔ has_deriv_at f (f' 1) x := +has_fderiv_at_filter_iff_has_deriv_at_filter + +lemma has_fderiv_at.has_deriv_at {f' : 𝕜 →L[𝕜] F} : + has_fderiv_at f f' x → has_deriv_at f (f' 1) x := +has_fderiv_at_iff_has_deriv_at.mp + +lemma has_strict_fderiv_at_iff_has_strict_deriv_at {f' : 𝕜 →L[𝕜] F} : + has_strict_fderiv_at f f' x ↔ has_strict_deriv_at f (f' 1) x := +by simp [has_strict_deriv_at, has_strict_fderiv_at] + +protected lemma has_strict_fderiv_at.has_strict_deriv_at {f' : 𝕜 →L[𝕜] F} : + has_strict_fderiv_at f f' x → has_strict_deriv_at f (f' 1) x := +has_strict_fderiv_at_iff_has_strict_deriv_at.mp + +lemma has_strict_deriv_at_iff_has_strict_fderiv_at : + has_strict_deriv_at f f' x ↔ has_strict_fderiv_at f (smul_right (1 : 𝕜 →L[𝕜] 𝕜) f') x := +iff.rfl + +alias has_strict_deriv_at_iff_has_strict_fderiv_at ↔ has_strict_deriv_at.has_strict_fderiv_at _ + +/-- Expressing `has_deriv_at f f' x` in terms of `has_fderiv_at` -/ +lemma has_deriv_at_iff_has_fderiv_at {f' : F} : + has_deriv_at f f' x ↔ + has_fderiv_at f (smul_right (1 : 𝕜 →L[𝕜] 𝕜) f') x := +iff.rfl + +alias has_deriv_at_iff_has_fderiv_at ↔ has_deriv_at.has_fderiv_at _ + +lemma deriv_within_zero_of_not_differentiable_within_at + (h : ¬ differentiable_within_at 𝕜 f s x) : deriv_within f s x = 0 := +by { unfold deriv_within, rw fderiv_within_zero_of_not_differentiable_within_at, simp, assumption } + +lemma differentiable_within_at_of_deriv_within_ne_zero (h : deriv_within f s x ≠ 0) : + differentiable_within_at 𝕜 f s x := +not_imp_comm.1 deriv_within_zero_of_not_differentiable_within_at h + +lemma deriv_zero_of_not_differentiable_at (h : ¬ differentiable_at 𝕜 f x) : deriv f x = 0 := +by { unfold deriv, rw fderiv_zero_of_not_differentiable_at, simp, assumption } + +lemma differentiable_at_of_deriv_ne_zero (h : deriv f x ≠ 0) : differentiable_at 𝕜 f x := +not_imp_comm.1 deriv_zero_of_not_differentiable_at h + +theorem unique_diff_within_at.eq_deriv (s : set 𝕜) (H : unique_diff_within_at 𝕜 s x) + (h : has_deriv_within_at f f' s x) (h₁ : has_deriv_within_at f f₁' s x) : f' = f₁' := +smul_right_one_eq_iff.mp $ unique_diff_within_at.eq H h h₁ + +theorem has_deriv_at_filter_iff_is_o : + has_deriv_at_filter f f' x L ↔ (λ x' : 𝕜, f x' - f x - (x' - x) • f') =o[L] (λ x', x' - x) := +iff.rfl + +theorem has_deriv_at_filter_iff_tendsto : + has_deriv_at_filter f f' x L ↔ + tendsto (λ x' : 𝕜, ‖x' - x‖⁻¹ * ‖f x' - f x - (x' - x) • f'‖) L (𝓝 0) := +has_fderiv_at_filter_iff_tendsto + +theorem has_deriv_within_at_iff_is_o : + has_deriv_within_at f f' s x + ↔ (λ x' : 𝕜, f x' - f x - (x' - x) • f') =o[𝓝[s] x] (λ x', x' - x) := +iff.rfl + +theorem has_deriv_within_at_iff_tendsto : has_deriv_within_at f f' s x ↔ + tendsto (λ x', ‖x' - x‖⁻¹ * ‖f x' - f x - (x' - x) • f'‖) (𝓝[s] x) (𝓝 0) := +has_fderiv_at_filter_iff_tendsto + +theorem has_deriv_at_iff_is_o : + has_deriv_at f f' x ↔ (λ x' : 𝕜, f x' - f x - (x' - x) • f') =o[𝓝 x] (λ x', x' - x) := +iff.rfl + +theorem has_deriv_at_iff_tendsto : has_deriv_at f f' x ↔ + tendsto (λ x', ‖x' - x‖⁻¹ * ‖f x' - f x - (x' - x) • f'‖) (𝓝 x) (𝓝 0) := +has_fderiv_at_filter_iff_tendsto + +theorem has_deriv_at_filter.is_O_sub (h : has_deriv_at_filter f f' x L) : + (λ x', f x' - f x) =O[L] (λ x', x' - x) := +has_fderiv_at_filter.is_O_sub h + +theorem has_deriv_at_filter.is_O_sub_rev (hf : has_deriv_at_filter f f' x L) (hf' : f' ≠ 0) : + (λ x', x' - x) =O[L] (λ x', f x' - f x) := +suffices antilipschitz_with ‖f'‖₊⁻¹ (smul_right (1 : 𝕜 →L[𝕜] 𝕜) f'), from hf.is_O_sub_rev this, +add_monoid_hom_class.antilipschitz_of_bound (smul_right (1 : 𝕜 →L[𝕜] 𝕜) f') $ + λ x, by simp [norm_smul, ← div_eq_inv_mul, mul_div_cancel _ (mt norm_eq_zero.1 hf')] + +theorem has_strict_deriv_at.has_deriv_at (h : has_strict_deriv_at f f' x) : + has_deriv_at f f' x := +h.has_fderiv_at +theorem has_deriv_within_at_congr_set' {s t : set 𝕜} (y : 𝕜) (h : s =ᶠ[𝓝[{y}ᶜ] x] t) : + has_deriv_within_at f f' s x ↔ has_deriv_within_at f f' t x := +has_fderiv_within_at_congr_set' y h + +theorem has_deriv_within_at_congr_set {s t : set 𝕜} (h : s =ᶠ[𝓝 x] t) : + has_deriv_within_at f f' s x ↔ has_deriv_within_at f f' t x := +has_fderiv_within_at_congr_set h + +alias has_deriv_within_at_congr_set ↔ has_deriv_within_at.congr_set _ + +@[simp] lemma has_deriv_within_at_diff_singleton : + has_deriv_within_at f f' (s \ {x}) x ↔ has_deriv_within_at f f' s x := +has_fderiv_within_at_diff_singleton _ + +@[simp] lemma has_deriv_within_at_Ioi_iff_Ici [partial_order 𝕜] : + has_deriv_within_at f f' (Ioi x) x ↔ has_deriv_within_at f f' (Ici x) x := +by rw [← Ici_diff_left, has_deriv_within_at_diff_singleton] + +alias has_deriv_within_at_Ioi_iff_Ici ↔ + has_deriv_within_at.Ici_of_Ioi has_deriv_within_at.Ioi_of_Ici + +@[simp] lemma has_deriv_within_at_Iio_iff_Iic [partial_order 𝕜] : + has_deriv_within_at f f' (Iio x) x ↔ has_deriv_within_at f f' (Iic x) x := +by rw [← Iic_diff_right, has_deriv_within_at_diff_singleton] + +alias has_deriv_within_at_Iio_iff_Iic ↔ + has_deriv_within_at.Iic_of_Iio has_deriv_within_at.Iio_of_Iic + +theorem has_deriv_within_at.Ioi_iff_Ioo [linear_order 𝕜] [order_closed_topology 𝕜] {x y : 𝕜} + (h : x < y) : + has_deriv_within_at f f' (Ioo x y) x ↔ has_deriv_within_at f f' (Ioi x) x := +has_fderiv_within_at_inter $ Iio_mem_nhds h + +alias has_deriv_within_at.Ioi_iff_Ioo ↔ + has_deriv_within_at.Ioi_of_Ioo has_deriv_within_at.Ioo_of_Ioi + +theorem has_deriv_at_iff_is_o_nhds_zero : has_deriv_at f f' x ↔ + (λh, f (x + h) - f x - h • f') =o[𝓝 0] (λh, h) := +has_fderiv_at_iff_is_o_nhds_zero + +theorem has_deriv_at_filter.mono (h : has_deriv_at_filter f f' x L₂) (hst : L₁ ≤ L₂) : + has_deriv_at_filter f f' x L₁ := +has_fderiv_at_filter.mono h hst + +theorem has_deriv_within_at.mono (h : has_deriv_within_at f f' t x) (hst : s ⊆ t) : + has_deriv_within_at f f' s x := +has_fderiv_within_at.mono h hst + +theorem has_deriv_within_at.mono_of_mem (h : has_deriv_within_at f f' t x) (hst : t ∈ 𝓝[s] x) : + has_deriv_within_at f f' s x := +has_fderiv_within_at.mono_of_mem h hst + +theorem has_deriv_at.has_deriv_at_filter (h : has_deriv_at f f' x) (hL : L ≤ 𝓝 x) : + has_deriv_at_filter f f' x L := +has_fderiv_at.has_fderiv_at_filter h hL + +theorem has_deriv_at.has_deriv_within_at + (h : has_deriv_at f f' x) : has_deriv_within_at f f' s x := +has_fderiv_at.has_fderiv_within_at h + +lemma has_deriv_within_at.differentiable_within_at (h : has_deriv_within_at f f' s x) : + differentiable_within_at 𝕜 f s x := +has_fderiv_within_at.differentiable_within_at h + +lemma has_deriv_at.differentiable_at (h : has_deriv_at f f' x) : differentiable_at 𝕜 f x := +has_fderiv_at.differentiable_at h + +@[simp] lemma has_deriv_within_at_univ : has_deriv_within_at f f' univ x ↔ has_deriv_at f f' x := +has_fderiv_within_at_univ + +theorem has_deriv_at.unique + (h₀ : has_deriv_at f f₀' x) (h₁ : has_deriv_at f f₁' x) : f₀' = f₁' := +smul_right_one_eq_iff.mp $ h₀.has_fderiv_at.unique h₁ + +lemma has_deriv_within_at_inter' (h : t ∈ 𝓝[s] x) : + has_deriv_within_at f f' (s ∩ t) x ↔ has_deriv_within_at f f' s x := +has_fderiv_within_at_inter' h + +lemma has_deriv_within_at_inter (h : t ∈ 𝓝 x) : + has_deriv_within_at f f' (s ∩ t) x ↔ has_deriv_within_at f f' s x := +has_fderiv_within_at_inter h + +lemma has_deriv_within_at.union (hs : has_deriv_within_at f f' s x) + (ht : has_deriv_within_at f f' t x) : + has_deriv_within_at f f' (s ∪ t) x := +hs.has_fderiv_within_at.union ht.has_fderiv_within_at + +lemma has_deriv_within_at.nhds_within (h : has_deriv_within_at f f' s x) + (ht : s ∈ 𝓝[t] x) : has_deriv_within_at f f' t x := +(has_deriv_within_at_inter' ht).1 (h.mono (inter_subset_right _ _)) + +lemma has_deriv_within_at.has_deriv_at (h : has_deriv_within_at f f' s x) (hs : s ∈ 𝓝 x) : + has_deriv_at f f' x := +has_fderiv_within_at.has_fderiv_at h hs + +lemma differentiable_within_at.has_deriv_within_at (h : differentiable_within_at 𝕜 f s x) : + has_deriv_within_at f (deriv_within f s x) s x := +h.has_fderiv_within_at.has_deriv_within_at + +lemma differentiable_at.has_deriv_at (h : differentiable_at 𝕜 f x) : has_deriv_at f (deriv f x) x := +h.has_fderiv_at.has_deriv_at + +@[simp] lemma has_deriv_at_deriv_iff : has_deriv_at f (deriv f x) x ↔ differentiable_at 𝕜 f x := +⟨λ h, h.differentiable_at, λ h, h.has_deriv_at⟩ + +@[simp] lemma has_deriv_within_at_deriv_within_iff : + has_deriv_within_at f (deriv_within f s x) s x ↔ differentiable_within_at 𝕜 f s x := +⟨λ h, h.differentiable_within_at, λ h, h.has_deriv_within_at⟩ + +lemma differentiable_on.has_deriv_at (h : differentiable_on 𝕜 f s) (hs : s ∈ 𝓝 x) : + has_deriv_at f (deriv f x) x := +(h.has_fderiv_at hs).has_deriv_at + +lemma has_deriv_at.deriv (h : has_deriv_at f f' x) : deriv f x = f' := +h.differentiable_at.has_deriv_at.unique h + +lemma deriv_eq {f' : 𝕜 → F} (h : ∀ x, has_deriv_at f (f' x) x) : deriv f = f' := +funext $ λ x, (h x).deriv + +lemma has_deriv_within_at.deriv_within + (h : has_deriv_within_at f f' s x) (hxs : unique_diff_within_at 𝕜 s x) : + deriv_within f s x = f' := +hxs.eq_deriv _ h.differentiable_within_at.has_deriv_within_at h + +lemma fderiv_within_deriv_within : (fderiv_within 𝕜 f s x : 𝕜 → F) 1 = deriv_within f s x := +rfl + +lemma deriv_within_fderiv_within : + smul_right (1 : 𝕜 →L[𝕜] 𝕜) (deriv_within f s x) = fderiv_within 𝕜 f s x := +by simp [deriv_within] + +lemma fderiv_deriv : (fderiv 𝕜 f x : 𝕜 → F) 1 = deriv f x := +rfl + +lemma deriv_fderiv : + smul_right (1 : 𝕜 →L[𝕜] 𝕜) (deriv f x) = fderiv 𝕜 f x := +by simp [deriv] + +lemma differentiable_at.deriv_within (h : differentiable_at 𝕜 f x) + (hxs : unique_diff_within_at 𝕜 s x) : deriv_within f s x = deriv f x := +by { unfold deriv_within deriv, rw h.fderiv_within hxs } + +theorem has_deriv_within_at.deriv_eq_zero (hd : has_deriv_within_at f 0 s x) + (H : unique_diff_within_at 𝕜 s x) : deriv f x = 0 := +(em' (differentiable_at 𝕜 f x)).elim deriv_zero_of_not_differentiable_at $ + λ h, H.eq_deriv _ h.has_deriv_at.has_deriv_within_at hd + +lemma deriv_within_of_mem (st : t ∈ 𝓝[s] x) (ht : unique_diff_within_at 𝕜 s x) + (h : differentiable_within_at 𝕜 f t x) : + deriv_within f s x = deriv_within f t x := +((differentiable_within_at.has_deriv_within_at h).mono_of_mem st).deriv_within ht + +lemma deriv_within_subset (st : s ⊆ t) (ht : unique_diff_within_at 𝕜 s x) + (h : differentiable_within_at 𝕜 f t x) : + deriv_within f s x = deriv_within f t x := +((differentiable_within_at.has_deriv_within_at h).mono st).deriv_within ht + +lemma deriv_within_congr_set' (y : 𝕜) (h : s =ᶠ[𝓝[{y}ᶜ] x] t) : + deriv_within f s x = deriv_within f t x := +by simp only [deriv_within, fderiv_within_congr_set' y h] + +lemma deriv_within_congr_set (h : s =ᶠ[𝓝 x] t) : deriv_within f s x = deriv_within f t x := +by simp only [deriv_within, fderiv_within_congr_set h] + +@[simp] lemma deriv_within_univ : deriv_within f univ = deriv f := +by { ext, unfold deriv_within deriv, rw fderiv_within_univ } + +lemma deriv_within_inter (ht : t ∈ 𝓝 x) : + deriv_within f (s ∩ t) x = deriv_within f s x := +by { unfold deriv_within, rw fderiv_within_inter ht } + +lemma deriv_within_of_open (hs : is_open s) (hx : x ∈ s) : + deriv_within f s x = deriv f x := +by { unfold deriv_within, rw fderiv_within_of_open hs hx, refl } + +lemma deriv_mem_iff {f : 𝕜 → F} {s : set F} {x : 𝕜} : + deriv f x ∈ s ↔ (differentiable_at 𝕜 f x ∧ deriv f x ∈ s) ∨ + (¬differentiable_at 𝕜 f x ∧ (0 : F) ∈ s) := +by by_cases hx : differentiable_at 𝕜 f x; simp [deriv_zero_of_not_differentiable_at, *] + +lemma deriv_within_mem_iff {f : 𝕜 → F} {t : set 𝕜} {s : set F} {x : 𝕜} : + deriv_within f t x ∈ s ↔ (differentiable_within_at 𝕜 f t x ∧ deriv_within f t x ∈ s) ∨ + (¬differentiable_within_at 𝕜 f t x ∧ (0 : F) ∈ s) := +by by_cases hx : differentiable_within_at 𝕜 f t x; + simp [deriv_within_zero_of_not_differentiable_within_at, *] + +lemma differentiable_within_at_Ioi_iff_Ici [partial_order 𝕜] : + differentiable_within_at 𝕜 f (Ioi x) x ↔ differentiable_within_at 𝕜 f (Ici x) x := +⟨λ h, h.has_deriv_within_at.Ici_of_Ioi.differentiable_within_at, +λ h, h.has_deriv_within_at.Ioi_of_Ici.differentiable_within_at⟩ + +-- Golfed while splitting the file +lemma deriv_within_Ioi_eq_Ici {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] (f : ℝ → E) + (x : ℝ) : + deriv_within f (Ioi x) x = deriv_within f (Ici x) x := +begin + by_cases H : differentiable_within_at ℝ f (Ioi x) x, + { have A := H.has_deriv_within_at.Ici_of_Ioi, + have B := (differentiable_within_at_Ioi_iff_Ici.1 H).has_deriv_within_at, + simpa using (unique_diff_on_Ici x).eq le_rfl A B }, + { rw [deriv_within_zero_of_not_differentiable_within_at H, + deriv_within_zero_of_not_differentiable_within_at], + rwa differentiable_within_at_Ioi_iff_Ici at H } +end + +section congr +/-! ### Congruence properties of derivatives -/ + +theorem filter.eventually_eq.has_deriv_at_filter_iff + (h₀ : f₀ =ᶠ[L] f₁) (hx : f₀ x = f₁ x) (h₁ : f₀' = f₁') : + has_deriv_at_filter f₀ f₀' x L ↔ has_deriv_at_filter f₁ f₁' x L := +h₀.has_fderiv_at_filter_iff hx (by simp [h₁]) + +lemma has_deriv_at_filter.congr_of_eventually_eq (h : has_deriv_at_filter f f' x L) + (hL : f₁ =ᶠ[L] f) (hx : f₁ x = f x) : has_deriv_at_filter f₁ f' x L := +by rwa hL.has_deriv_at_filter_iff hx rfl + +lemma has_deriv_within_at.congr_mono (h : has_deriv_within_at f f' s x) (ht : ∀x ∈ t, f₁ x = f x) + (hx : f₁ x = f x) (h₁ : t ⊆ s) : has_deriv_within_at f₁ f' t x := +has_fderiv_within_at.congr_mono h ht hx h₁ + +lemma has_deriv_within_at.congr (h : has_deriv_within_at f f' s x) (hs : ∀x ∈ s, f₁ x = f x) + (hx : f₁ x = f x) : has_deriv_within_at f₁ f' s x := +h.congr_mono hs hx (subset.refl _) + +lemma has_deriv_within_at.congr_of_mem (h : has_deriv_within_at f f' s x) (hs : ∀x ∈ s, f₁ x = f x) + (hx : x ∈ s) : has_deriv_within_at f₁ f' s x := +h.congr hs (hs _ hx) + +lemma has_deriv_within_at.congr_of_eventually_eq (h : has_deriv_within_at f f' s x) + (h₁ : f₁ =ᶠ[𝓝[s] x] f) (hx : f₁ x = f x) : has_deriv_within_at f₁ f' s x := +has_deriv_at_filter.congr_of_eventually_eq h h₁ hx + +lemma has_deriv_within_at.congr_of_eventually_eq_of_mem (h : has_deriv_within_at f f' s x) + (h₁ : f₁ =ᶠ[𝓝[s] x] f) (hx : x ∈ s) : has_deriv_within_at f₁ f' s x := +h.congr_of_eventually_eq h₁ (h₁.eq_of_nhds_within hx) + +lemma has_deriv_at.congr_of_eventually_eq (h : has_deriv_at f f' x) + (h₁ : f₁ =ᶠ[𝓝 x] f) : has_deriv_at f₁ f' x := +has_deriv_at_filter.congr_of_eventually_eq h h₁ (mem_of_mem_nhds h₁ : _) + +lemma filter.eventually_eq.deriv_within_eq (hL : f₁ =ᶠ[𝓝[s] x] f) (hx : f₁ x = f x) : + deriv_within f₁ s x = deriv_within f s x := +by { unfold deriv_within, rw hL.fderiv_within_eq hx } + +lemma deriv_within_congr (hs : eq_on f₁ f s) (hx : f₁ x = f x) : + deriv_within f₁ s x = deriv_within f s x := +by { unfold deriv_within, rw fderiv_within_congr hs hx } + +lemma filter.eventually_eq.deriv_eq (hL : f₁ =ᶠ[𝓝 x] f) : deriv f₁ x = deriv f x := +by { unfold deriv, rwa filter.eventually_eq.fderiv_eq } + +protected lemma filter.eventually_eq.deriv (h : f₁ =ᶠ[𝓝 x] f) : deriv f₁ =ᶠ[𝓝 x] deriv f := +h.eventually_eq_nhds.mono $ λ x h, h.deriv_eq + +end congr + +section id +/-! ### Derivative of the identity -/ +variables (s x L) + +theorem has_deriv_at_filter_id : has_deriv_at_filter id 1 x L := +(has_fderiv_at_filter_id x L).has_deriv_at_filter + +theorem has_deriv_within_at_id : has_deriv_within_at id 1 s x := +has_deriv_at_filter_id _ _ + +theorem has_deriv_at_id : has_deriv_at id 1 x := +has_deriv_at_filter_id _ _ + +theorem has_deriv_at_id' : has_deriv_at (λ (x : 𝕜), x) 1 x := +has_deriv_at_filter_id _ _ + +theorem has_strict_deriv_at_id : has_strict_deriv_at id 1 x := +(has_strict_fderiv_at_id x).has_strict_deriv_at + +lemma deriv_id : deriv id x = 1 := +has_deriv_at.deriv (has_deriv_at_id x) + +@[simp] lemma deriv_id' : deriv (@id 𝕜) = λ _, 1 := funext deriv_id + +@[simp] lemma deriv_id'' : deriv (λ x : 𝕜, x) = λ _, 1 := deriv_id' + +lemma deriv_within_id (hxs : unique_diff_within_at 𝕜 s x) : deriv_within id s x = 1 := +(has_deriv_within_at_id x s).deriv_within hxs + +end id + +section const +/-! ### Derivative of constant functions -/ +variables (c : F) (s x L) + +theorem has_deriv_at_filter_const : has_deriv_at_filter (λ x, c) 0 x L := +(has_fderiv_at_filter_const c x L).has_deriv_at_filter + +theorem has_strict_deriv_at_const : has_strict_deriv_at (λ x, c) 0 x := +(has_strict_fderiv_at_const c x).has_strict_deriv_at + +theorem has_deriv_within_at_const : has_deriv_within_at (λ x, c) 0 s x := +has_deriv_at_filter_const _ _ _ + +theorem has_deriv_at_const : has_deriv_at (λ x, c) 0 x := +has_deriv_at_filter_const _ _ _ + +lemma deriv_const : deriv (λ x, c) x = 0 := +has_deriv_at.deriv (has_deriv_at_const x c) + +@[simp] lemma deriv_const' : deriv (λ x:𝕜, c) = λ x, 0 := +funext (λ x, deriv_const x c) + +lemma deriv_within_const (hxs : unique_diff_within_at 𝕜 s x) : deriv_within (λ x, c) s x = 0 := +(has_deriv_within_at_const _ _ _).deriv_within hxs + +end const + +section continuous +/-! ### Continuity of a function admitting a derivative -/ + +theorem has_deriv_at_filter.tendsto_nhds + (hL : L ≤ 𝓝 x) (h : has_deriv_at_filter f f' x L) : + tendsto f L (𝓝 (f x)) := +h.tendsto_nhds hL + +theorem has_deriv_within_at.continuous_within_at + (h : has_deriv_within_at f f' s x) : continuous_within_at f s x := +has_deriv_at_filter.tendsto_nhds inf_le_left h + +theorem has_deriv_at.continuous_at (h : has_deriv_at f f' x) : continuous_at f x := +has_deriv_at_filter.tendsto_nhds le_rfl h + +protected theorem has_deriv_at.continuous_on {f f' : 𝕜 → F} + (hderiv : ∀ x ∈ s, has_deriv_at f (f' x) x) : continuous_on f s := +λ x hx, (hderiv x hx).continuous_at.continuous_within_at + +end continuous + diff --git a/src/analysis/calculus/deriv/comp.lean b/src/analysis/calculus/deriv/comp.lean new file mode 100644 index 0000000000000..1fdb3fd794b61 --- /dev/null +++ b/src/analysis/calculus/deriv/comp.lean @@ -0,0 +1,263 @@ +/- +Copyright (c) 2019 Gabriel Ebner. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Gabriel Ebner, Sébastien Gouëzel, Yury Kudryashov, Yuyang Zhao +-/ +import analysis.calculus.deriv.basic +import analysis.calculus.fderiv.comp +import analysis.calculus.fderiv.restrict_scalars + +/-! +# One-dimensional derivatives of compositions of functions + +In this file we prove the chain rule for the following cases: + +* `has_deriv_at.comp` etc: `f : 𝕜' → 𝕜'` composed with `g : 𝕜 → 𝕜'`; +* `has_deriv_at.scomp` etc: `f : 𝕜' → E` composed with `g : 𝕜 → 𝕜'`; +* `has_fderiv_at.comp_has_deriv_at` etc: `f : E → F` composed with `g : 𝕜 → E`; + +Here `𝕜` is the base normed field, `E` and `F` are normed spaces over `𝕜` and `𝕜'` is an algebra +over `𝕜` (e.g., `𝕜'=𝕜` or `𝕜=ℝ`, `𝕜'=ℂ`). + +For a more detailed overview of one-dimensional derivatives in mathlib, see the module docstring of +`analysis/calculus/deriv/basic`. + +## Keywords + +derivative, chain rule +-/ + +universes u v w +open_locale classical topology big_operators filter ennreal +open filter asymptotics set +open continuous_linear_map (smul_right smul_right_one_eq_iff) + +variables {𝕜 : Type u} [nontrivially_normed_field 𝕜] +variables {F : Type v} [normed_add_comm_group F] [normed_space 𝕜 F] +variables {E : Type w} [normed_add_comm_group E] [normed_space 𝕜 E] + +variables {f f₀ f₁ g : 𝕜 → F} +variables {f' f₀' f₁' g' : F} +variables {x : 𝕜} +variables {s t : set 𝕜} +variables {L L₁ L₂ : filter 𝕜} + +section composition +/-! +### Derivative of the composition of a vector function and a scalar function + +We use `scomp` in lemmas on composition of vector valued and scalar valued functions, and `comp` +in lemmas on composition of scalar valued functions, in analogy for `smul` and `mul` (and also +because the `comp` version with the shorter name will show up much more often in applications). +The formula for the derivative involves `smul` in `scomp` lemmas, which can be reduced to +usual multiplication in `comp` lemmas. +-/ + +/- For composition lemmas, we put x explicit to help the elaborator, as otherwise Lean tends to +get confused since there are too many possibilities for composition -/ +variables {𝕜' : Type*} [nontrivially_normed_field 𝕜'] [normed_algebra 𝕜 𝕜'] + [normed_space 𝕜' F] [is_scalar_tower 𝕜 𝕜' F] {s' t' : set 𝕜'} + {h : 𝕜 → 𝕜'} {h₁ : 𝕜 → 𝕜} {h₂ : 𝕜' → 𝕜'} {h' h₂' : 𝕜'} {h₁' : 𝕜} + {g₁ : 𝕜' → F} {g₁' : F} {L' : filter 𝕜'} (x) + +theorem has_deriv_at_filter.scomp + (hg : has_deriv_at_filter g₁ g₁' (h x) L') + (hh : has_deriv_at_filter h h' x L) (hL : tendsto h L L'): + has_deriv_at_filter (g₁ ∘ h) (h' • g₁') x L := +by simpa using ((hg.restrict_scalars 𝕜).comp x hh hL).has_deriv_at_filter + +theorem has_deriv_within_at.scomp_has_deriv_at + (hg : has_deriv_within_at g₁ g₁' s' (h x)) + (hh : has_deriv_at h h' x) (hs : ∀ x, h x ∈ s') : + has_deriv_at (g₁ ∘ h) (h' • g₁') x := +hg.scomp x hh $ tendsto_inf.2 ⟨hh.continuous_at, tendsto_principal.2 $ eventually_of_forall hs⟩ + +theorem has_deriv_within_at.scomp + (hg : has_deriv_within_at g₁ g₁' t' (h x)) + (hh : has_deriv_within_at h h' s x) (hst : maps_to h s t') : + has_deriv_within_at (g₁ ∘ h) (h' • g₁') s x := +hg.scomp x hh $ hh.continuous_within_at.tendsto_nhds_within hst + +/-- The chain rule. -/ +theorem has_deriv_at.scomp + (hg : has_deriv_at g₁ g₁' (h x)) (hh : has_deriv_at h h' x) : + has_deriv_at (g₁ ∘ h) (h' • g₁') x := +hg.scomp x hh hh.continuous_at + +theorem has_strict_deriv_at.scomp + (hg : has_strict_deriv_at g₁ g₁' (h x)) (hh : has_strict_deriv_at h h' x) : + has_strict_deriv_at (g₁ ∘ h) (h' • g₁') x := +by simpa using ((hg.restrict_scalars 𝕜).comp x hh).has_strict_deriv_at + +theorem has_deriv_at.scomp_has_deriv_within_at + (hg : has_deriv_at g₁ g₁' (h x)) (hh : has_deriv_within_at h h' s x) : + has_deriv_within_at (g₁ ∘ h) (h' • g₁') s x := +has_deriv_within_at.scomp x hg.has_deriv_within_at hh (maps_to_univ _ _) + +lemma deriv_within.scomp + (hg : differentiable_within_at 𝕜' g₁ t' (h x)) (hh : differentiable_within_at 𝕜 h s x) + (hs : maps_to h s t') (hxs : unique_diff_within_at 𝕜 s x) : + deriv_within (g₁ ∘ h) s x = deriv_within h s x • deriv_within g₁ t' (h x) := +(has_deriv_within_at.scomp x hg.has_deriv_within_at hh.has_deriv_within_at hs).deriv_within hxs + +lemma deriv.scomp + (hg : differentiable_at 𝕜' g₁ (h x)) (hh : differentiable_at 𝕜 h x) : + deriv (g₁ ∘ h) x = deriv h x • deriv g₁ (h x) := +(has_deriv_at.scomp x hg.has_deriv_at hh.has_deriv_at).deriv + +/-! ### Derivative of the composition of a scalar and vector functions -/ + +theorem has_deriv_at_filter.comp_has_fderiv_at_filter {f : E → 𝕜'} {f' : E →L[𝕜] 𝕜'} (x) + {L'' : filter E} (hh₂ : has_deriv_at_filter h₂ h₂' (f x) L') + (hf : has_fderiv_at_filter f f' x L'') (hL : tendsto f L'' L') : + has_fderiv_at_filter (h₂ ∘ f) (h₂' • f') x L'' := +by { convert (hh₂.restrict_scalars 𝕜).comp x hf hL, ext x, simp [mul_comm] } + +theorem has_strict_deriv_at.comp_has_strict_fderiv_at {f : E → 𝕜'} {f' : E →L[𝕜] 𝕜'} (x) + (hh : has_strict_deriv_at h₂ h₂' (f x)) (hf : has_strict_fderiv_at f f' x) : + has_strict_fderiv_at (h₂ ∘ f) (h₂' • f') x := +begin + rw has_strict_deriv_at at hh, + convert (hh.restrict_scalars 𝕜).comp x hf, + ext x, + simp [mul_comm] +end + +theorem has_deriv_at.comp_has_fderiv_at {f : E → 𝕜'} {f' : E →L[𝕜] 𝕜'} (x) + (hh : has_deriv_at h₂ h₂' (f x)) (hf : has_fderiv_at f f' x) : + has_fderiv_at (h₂ ∘ f) (h₂' • f') x := +hh.comp_has_fderiv_at_filter x hf hf.continuous_at + +theorem has_deriv_at.comp_has_fderiv_within_at {f : E → 𝕜'} {f' : E →L[𝕜] 𝕜'} {s} (x) + (hh : has_deriv_at h₂ h₂' (f x)) (hf : has_fderiv_within_at f f' s x) : + has_fderiv_within_at (h₂ ∘ f) (h₂' • f') s x := +hh.comp_has_fderiv_at_filter x hf hf.continuous_within_at + +theorem has_deriv_within_at.comp_has_fderiv_within_at {f : E → 𝕜'} {f' : E →L[𝕜] 𝕜'} {s t} (x) + (hh : has_deriv_within_at h₂ h₂' t (f x)) (hf : has_fderiv_within_at f f' s x) + (hst : maps_to f s t) : + has_fderiv_within_at (h₂ ∘ f) (h₂' • f') s x := +hh.comp_has_fderiv_at_filter x hf $ hf.continuous_within_at.tendsto_nhds_within hst + +/-! ### Derivative of the composition of two scalar functions -/ + +theorem has_deriv_at_filter.comp + (hh₂ : has_deriv_at_filter h₂ h₂' (h x) L') + (hh : has_deriv_at_filter h h' x L) (hL : tendsto h L L') : + has_deriv_at_filter (h₂ ∘ h) (h₂' * h') x L := +by { rw mul_comm, exact hh₂.scomp x hh hL } + +theorem has_deriv_within_at.comp + (hh₂ : has_deriv_within_at h₂ h₂' s' (h x)) + (hh : has_deriv_within_at h h' s x) (hst : maps_to h s s') : + has_deriv_within_at (h₂ ∘ h) (h₂' * h') s x := +by { rw mul_comm, exact hh₂.scomp x hh hst, } + +/-- The chain rule. -/ +theorem has_deriv_at.comp + (hh₂ : has_deriv_at h₂ h₂' (h x)) (hh : has_deriv_at h h' x) : + has_deriv_at (h₂ ∘ h) (h₂' * h') x := +hh₂.comp x hh hh.continuous_at + +theorem has_strict_deriv_at.comp + (hh₂ : has_strict_deriv_at h₂ h₂' (h x)) (hh : has_strict_deriv_at h h' x) : + has_strict_deriv_at (h₂ ∘ h) (h₂' * h') x := +by { rw mul_comm, exact hh₂.scomp x hh } + +theorem has_deriv_at.comp_has_deriv_within_at + (hh₂ : has_deriv_at h₂ h₂' (h x)) (hh : has_deriv_within_at h h' s x) : + has_deriv_within_at (h₂ ∘ h) (h₂' * h') s x := +hh₂.has_deriv_within_at.comp x hh (maps_to_univ _ _) + +lemma deriv_within.comp + (hh₂ : differentiable_within_at 𝕜' h₂ s' (h x)) (hh : differentiable_within_at 𝕜 h s x) + (hs : maps_to h s s') (hxs : unique_diff_within_at 𝕜 s x) : + deriv_within (h₂ ∘ h) s x = deriv_within h₂ s' (h x) * deriv_within h s x := +(hh₂.has_deriv_within_at.comp x hh.has_deriv_within_at hs).deriv_within hxs + +lemma deriv.comp + (hh₂ : differentiable_at 𝕜' h₂ (h x)) (hh : differentiable_at 𝕜 h x) : + deriv (h₂ ∘ h) x = deriv h₂ (h x) * deriv h x := +(hh₂.has_deriv_at.comp x hh.has_deriv_at).deriv + +protected lemma has_deriv_at_filter.iterate {f : 𝕜 → 𝕜} {f' : 𝕜} + (hf : has_deriv_at_filter f f' x L) (hL : tendsto f L L) (hx : f x = x) (n : ℕ) : + has_deriv_at_filter (f^[n]) (f'^n) x L := +begin + have := hf.iterate hL hx n, + rwa [continuous_linear_map.smul_right_one_pow] at this +end + +protected lemma has_deriv_at.iterate {f : 𝕜 → 𝕜} {f' : 𝕜} + (hf : has_deriv_at f f' x) (hx : f x = x) (n : ℕ) : + has_deriv_at (f^[n]) (f'^n) x := +begin + have := has_fderiv_at.iterate hf hx n, + rwa [continuous_linear_map.smul_right_one_pow] at this +end + +protected lemma has_deriv_within_at.iterate {f : 𝕜 → 𝕜} {f' : 𝕜} + (hf : has_deriv_within_at f f' s x) (hx : f x = x) (hs : maps_to f s s) (n : ℕ) : + has_deriv_within_at (f^[n]) (f'^n) s x := +begin + have := has_fderiv_within_at.iterate hf hx hs n, + rwa [continuous_linear_map.smul_right_one_pow] at this +end + +protected lemma has_strict_deriv_at.iterate {f : 𝕜 → 𝕜} {f' : 𝕜} + (hf : has_strict_deriv_at f f' x) (hx : f x = x) (n : ℕ) : + has_strict_deriv_at (f^[n]) (f'^n) x := +begin + have := hf.iterate hx n, + rwa [continuous_linear_map.smul_right_one_pow] at this +end + +end composition + +section composition_vector +/-! ### Derivative of the composition of a function between vector spaces and a function on `𝕜` -/ + +open continuous_linear_map + +variables {l : F → E} {l' : F →L[𝕜] E} +variable (x) + +/-- The composition `l ∘ f` where `l : F → E` and `f : 𝕜 → F`, has a derivative within a set +equal to the Fréchet derivative of `l` applied to the derivative of `f`. -/ +theorem has_fderiv_within_at.comp_has_deriv_within_at {t : set F} + (hl : has_fderiv_within_at l l' t (f x)) (hf : has_deriv_within_at f f' s x) + (hst : maps_to f s t) : + has_deriv_within_at (l ∘ f) (l' f') s x := +by simpa only [one_apply, one_smul, smul_right_apply, coe_comp', (∘)] + using (hl.comp x hf.has_fderiv_within_at hst).has_deriv_within_at + +theorem has_fderiv_at.comp_has_deriv_within_at + (hl : has_fderiv_at l l' (f x)) (hf : has_deriv_within_at f f' s x) : + has_deriv_within_at (l ∘ f) (l' f') s x := +hl.has_fderiv_within_at.comp_has_deriv_within_at x hf (maps_to_univ _ _) + +/-- The composition `l ∘ f` where `l : F → E` and `f : 𝕜 → F`, has a derivative equal to the +Fréchet derivative of `l` applied to the derivative of `f`. -/ +theorem has_fderiv_at.comp_has_deriv_at (hl : has_fderiv_at l l' (f x)) (hf : has_deriv_at f f' x) : + has_deriv_at (l ∘ f) (l' f') x := +has_deriv_within_at_univ.mp $ hl.comp_has_deriv_within_at x hf.has_deriv_within_at + +theorem has_strict_fderiv_at.comp_has_strict_deriv_at + (hl : has_strict_fderiv_at l l' (f x)) (hf : has_strict_deriv_at f f' x) : + has_strict_deriv_at (l ∘ f) (l' f') x := +by simpa only [one_apply, one_smul, smul_right_apply, coe_comp', (∘)] + using (hl.comp x hf.has_strict_fderiv_at).has_strict_deriv_at + +lemma fderiv_within.comp_deriv_within {t : set F} + (hl : differentiable_within_at 𝕜 l t (f x)) (hf : differentiable_within_at 𝕜 f s x) + (hs : maps_to f s t) (hxs : unique_diff_within_at 𝕜 s x) : + deriv_within (l ∘ f) s x = (fderiv_within 𝕜 l t (f x) : F → E) (deriv_within f s x) := +(hl.has_fderiv_within_at.comp_has_deriv_within_at x hf.has_deriv_within_at hs).deriv_within hxs + +lemma fderiv.comp_deriv + (hl : differentiable_at 𝕜 l (f x)) (hf : differentiable_at 𝕜 f x) : + deriv (l ∘ f) x = (fderiv 𝕜 l (f x) : F → E) (deriv f x) := +(hl.has_fderiv_at.comp_has_deriv_at x hf.has_deriv_at).deriv + +end composition_vector + diff --git a/src/analysis/calculus/deriv/inv.lean b/src/analysis/calculus/deriv/inv.lean new file mode 100644 index 0000000000000..ec37caff033d7 --- /dev/null +++ b/src/analysis/calculus/deriv/inv.lean @@ -0,0 +1,218 @@ +/- +Copyright (c) 2023 Sébastien Gouëzel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sébastien Gouëzel, Yury Kudryashov +-/ +import analysis.calculus.deriv.mul +import analysis.calculus.deriv.comp + +/-! +# Derivatives of `x ↦ x⁻¹` and `f x / g x` + +In this file we prove `(x⁻¹)' = -1 / x ^ 2`, `((f x)⁻¹)' = -f' x / (f x) ^ 2`, and +`(f x / g x)' = (f' x * g x - f x * g' x) / (g x) ^ 2` for different notions of derivative. + +For a more detailed overview of one-dimensional derivatives in mathlib, see the module docstring of +`analysis/calculus/deriv/basic`. + +## Keywords + +derivative +-/ + +universes u v w +open_locale classical topology big_operators filter ennreal +open filter asymptotics set +open continuous_linear_map (smul_right smul_right_one_eq_iff) + + +variables {𝕜 : Type u} [nontrivially_normed_field 𝕜] +variables {F : Type v} [normed_add_comm_group F] [normed_space 𝕜 F] +variables {E : Type w} [normed_add_comm_group E] [normed_space 𝕜 E] + +variables {f f₀ f₁ g : 𝕜 → F} +variables {f' f₀' f₁' g' : F} +variables {x : 𝕜} +variables {s t : set 𝕜} +variables {L : filter 𝕜} + +section inverse +/-! ### Derivative of `x ↦ x⁻¹` -/ + +theorem has_strict_deriv_at_inv (hx : x ≠ 0) : has_strict_deriv_at has_inv.inv (-(x^2)⁻¹) x := +begin + suffices : (λ p : 𝕜 × 𝕜, (p.1 - p.2) * ((x * x)⁻¹ - (p.1 * p.2)⁻¹)) =o[𝓝 (x, x)] + (λ p, (p.1 - p.2) * 1), + { refine this.congr' _ (eventually_of_forall $ λ _, mul_one _), + refine eventually.mono ((is_open_ne.prod is_open_ne).mem_nhds ⟨hx, hx⟩) _, + rintro ⟨y, z⟩ ⟨hy, hz⟩, + simp only [mem_set_of_eq] at hy hz, -- hy : y ≠ 0, hz : z ≠ 0 + field_simp [hx, hy, hz], ring, }, + refine (is_O_refl (λ p : 𝕜 × 𝕜, p.1 - p.2) _).mul_is_o ((is_o_one_iff _).2 _), + rw [← sub_self (x * x)⁻¹], + exact tendsto_const_nhds.sub ((continuous_mul.tendsto (x, x)).inv₀ $ mul_ne_zero hx hx) +end + +theorem has_deriv_at_inv (x_ne_zero : x ≠ 0) : + has_deriv_at (λy, y⁻¹) (-(x^2)⁻¹) x := +(has_strict_deriv_at_inv x_ne_zero).has_deriv_at + +theorem has_deriv_within_at_inv (x_ne_zero : x ≠ 0) (s : set 𝕜) : + has_deriv_within_at (λx, x⁻¹) (-(x^2)⁻¹) s x := +(has_deriv_at_inv x_ne_zero).has_deriv_within_at + +lemma differentiable_at_inv : + differentiable_at 𝕜 (λx, x⁻¹) x ↔ x ≠ 0:= +⟨λ H, normed_field.continuous_at_inv.1 H.continuous_at, + λ H, (has_deriv_at_inv H).differentiable_at⟩ + +lemma differentiable_within_at_inv (x_ne_zero : x ≠ 0) : + differentiable_within_at 𝕜 (λx, x⁻¹) s x := +(differentiable_at_inv.2 x_ne_zero).differentiable_within_at + +lemma differentiable_on_inv : differentiable_on 𝕜 (λx:𝕜, x⁻¹) {x | x ≠ 0} := +λx hx, differentiable_within_at_inv hx + +lemma deriv_inv : deriv (λx, x⁻¹) x = -(x^2)⁻¹ := +begin + rcases eq_or_ne x 0 with rfl|hne, + { simp [deriv_zero_of_not_differentiable_at (mt differentiable_at_inv.1 (not_not.2 rfl))] }, + { exact (has_deriv_at_inv hne).deriv } +end + +@[simp] lemma deriv_inv' : deriv (λ x : 𝕜, x⁻¹) = λ x, -(x ^ 2)⁻¹ := funext (λ x, deriv_inv) + +lemma deriv_within_inv (x_ne_zero : x ≠ 0) (hxs : unique_diff_within_at 𝕜 s x) : + deriv_within (λx, x⁻¹) s x = -(x^2)⁻¹ := +begin + rw differentiable_at.deriv_within (differentiable_at_inv.2 x_ne_zero) hxs, + exact deriv_inv +end + +lemma has_fderiv_at_inv (x_ne_zero : x ≠ 0) : + has_fderiv_at (λx, x⁻¹) (smul_right (1 : 𝕜 →L[𝕜] 𝕜) (-(x^2)⁻¹) : 𝕜 →L[𝕜] 𝕜) x := +has_deriv_at_inv x_ne_zero + +lemma has_fderiv_within_at_inv (x_ne_zero : x ≠ 0) : + has_fderiv_within_at (λx, x⁻¹) (smul_right (1 : 𝕜 →L[𝕜] 𝕜) (-(x^2)⁻¹) : 𝕜 →L[𝕜] 𝕜) s x := +(has_fderiv_at_inv x_ne_zero).has_fderiv_within_at + +lemma fderiv_inv : + fderiv 𝕜 (λx, x⁻¹) x = smul_right (1 : 𝕜 →L[𝕜] 𝕜) (-(x^2)⁻¹) := +by rw [← deriv_fderiv, deriv_inv] + +lemma fderiv_within_inv (x_ne_zero : x ≠ 0) (hxs : unique_diff_within_at 𝕜 s x) : + fderiv_within 𝕜 (λx, x⁻¹) s x = smul_right (1 : 𝕜 →L[𝕜] 𝕜) (-(x^2)⁻¹) := +begin + rw differentiable_at.fderiv_within (differentiable_at_inv.2 x_ne_zero) hxs, + exact fderiv_inv +end + +variables {c : 𝕜 → 𝕜} {h : E → 𝕜} {c' : 𝕜} {z : E} {S : set E} + +lemma has_deriv_within_at.inv + (hc : has_deriv_within_at c c' s x) (hx : c x ≠ 0) : + has_deriv_within_at (λ y, (c y)⁻¹) (- c' / (c x)^2) s x := +begin + convert (has_deriv_at_inv hx).comp_has_deriv_within_at x hc, + field_simp +end + +lemma has_deriv_at.inv (hc : has_deriv_at c c' x) (hx : c x ≠ 0) : + has_deriv_at (λ y, (c y)⁻¹) (- c' / (c x)^2) x := +begin + rw ← has_deriv_within_at_univ at *, + exact hc.inv hx +end + +lemma differentiable_within_at.inv (hf : differentiable_within_at 𝕜 h S z) (hz : h z ≠ 0) : + differentiable_within_at 𝕜 (λx, (h x)⁻¹) S z := +(differentiable_at_inv.mpr hz).comp_differentiable_within_at z hf + +@[simp] lemma differentiable_at.inv (hf : differentiable_at 𝕜 h z) (hz : h z ≠ 0) : + differentiable_at 𝕜 (λx, (h x)⁻¹) z := +(differentiable_at_inv.mpr hz).comp z hf + +lemma differentiable_on.inv (hf : differentiable_on 𝕜 h S) (hz : ∀ x ∈ S, h x ≠ 0) : + differentiable_on 𝕜 (λx, (h x)⁻¹) S := +λx h, (hf x h).inv (hz x h) + +@[simp] lemma differentiable.inv (hf : differentiable 𝕜 h) (hz : ∀ x, h x ≠ 0) : + differentiable 𝕜 (λx, (h x)⁻¹) := +λx, (hf x).inv (hz x) + +lemma deriv_within_inv' (hc : differentiable_within_at 𝕜 c s x) (hx : c x ≠ 0) + (hxs : unique_diff_within_at 𝕜 s x) : + deriv_within (λx, (c x)⁻¹) s x = - (deriv_within c s x) / (c x)^2 := +(hc.has_deriv_within_at.inv hx).deriv_within hxs + +@[simp] lemma deriv_inv'' (hc : differentiable_at 𝕜 c x) (hx : c x ≠ 0) : + deriv (λx, (c x)⁻¹) x = - (deriv c x) / (c x)^2 := +(hc.has_deriv_at.inv hx).deriv + +end inverse + +section division +/-! ### Derivative of `x ↦ c x / d x` -/ + +variables {𝕜' : Type*} [nontrivially_normed_field 𝕜'] [normed_algebra 𝕜 𝕜'] + {c d : 𝕜 → 𝕜'} {c' d' : 𝕜'} + +lemma has_deriv_within_at.div + (hc : has_deriv_within_at c c' s x) (hd : has_deriv_within_at d d' s x) (hx : d x ≠ 0) : + has_deriv_within_at (λ y, c y / d y) ((c' * d x - c x * d') / (d x)^2) s x := +begin + convert hc.mul ((has_deriv_at_inv hx).comp_has_deriv_within_at x hd), + { simp only [div_eq_mul_inv] }, + { field_simp, ring } +end + +lemma has_strict_deriv_at.div (hc : has_strict_deriv_at c c' x) (hd : has_strict_deriv_at d d' x) + (hx : d x ≠ 0) : + has_strict_deriv_at (λ y, c y / d y) ((c' * d x - c x * d') / (d x)^2) x := +begin + convert hc.mul ((has_strict_deriv_at_inv hx).comp x hd), + { simp only [div_eq_mul_inv] }, + { field_simp, ring } +end + +lemma has_deriv_at.div (hc : has_deriv_at c c' x) (hd : has_deriv_at d d' x) (hx : d x ≠ 0) : + has_deriv_at (λ y, c y / d y) ((c' * d x - c x * d') / (d x)^2) x := +begin + rw ← has_deriv_within_at_univ at *, + exact hc.div hd hx +end + +lemma differentiable_within_at.div + (hc : differentiable_within_at 𝕜 c s x) (hd : differentiable_within_at 𝕜 d s x) (hx : d x ≠ 0) : + differentiable_within_at 𝕜 (λx, c x / d x) s x := +((hc.has_deriv_within_at).div (hd.has_deriv_within_at) hx).differentiable_within_at + +@[simp] lemma differentiable_at.div + (hc : differentiable_at 𝕜 c x) (hd : differentiable_at 𝕜 d x) (hx : d x ≠ 0) : + differentiable_at 𝕜 (λx, c x / d x) x := +((hc.has_deriv_at).div (hd.has_deriv_at) hx).differentiable_at + +lemma differentiable_on.div + (hc : differentiable_on 𝕜 c s) (hd : differentiable_on 𝕜 d s) (hx : ∀ x ∈ s, d x ≠ 0) : + differentiable_on 𝕜 (λx, c x / d x) s := +λx h, (hc x h).div (hd x h) (hx x h) + +@[simp] lemma differentiable.div + (hc : differentiable 𝕜 c) (hd : differentiable 𝕜 d) (hx : ∀ x, d x ≠ 0) : +differentiable 𝕜 (λx, c x / d x) := +λx, (hc x).div (hd x) (hx x) + +lemma deriv_within_div + (hc : differentiable_within_at 𝕜 c s x) (hd : differentiable_within_at 𝕜 d s x) (hx : d x ≠ 0) + (hxs : unique_diff_within_at 𝕜 s x) : + deriv_within (λx, c x / d x) s x + = ((deriv_within c s x) * d x - c x * (deriv_within d s x)) / (d x)^2 := +((hc.has_deriv_within_at).div (hd.has_deriv_within_at) hx).deriv_within hxs + +@[simp] lemma deriv_div + (hc : differentiable_at 𝕜 c x) (hd : differentiable_at 𝕜 d x) (hx : d x ≠ 0) : + deriv (λx, c x / d x) x = ((deriv c x) * d x - c x * (deriv d x)) / (d x)^2 := +((hc.has_deriv_at).div (hd.has_deriv_at) hx).deriv + +end division diff --git a/src/analysis/calculus/deriv/inverse.lean b/src/analysis/calculus/deriv/inverse.lean new file mode 100644 index 0000000000000..4e21ac6ced7f9 --- /dev/null +++ b/src/analysis/calculus/deriv/inverse.lean @@ -0,0 +1,120 @@ +/- +Copyright (c) 2021 Yury Kudryashov. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yury Kudryashov +-/ +import analysis.calculus.deriv.comp +import analysis.calculus.fderiv.equiv + +/-! +# Inverse function theorem - the easy half + +In this file we prove that `g' (f x) = (f' x)⁻¹` provided that `f` is strictly differentiable at +`x`, `f' x ≠ 0`, and `g` is a local left inverse of `f` that is continuous at `f x`. This is the +easy half of the inverse function theorem: the harder half states that `g` exists. + +For a more detailed overview of one-dimensional derivatives in mathlib, see the module docstring of +`analysis/calculus/deriv/basic`. + +## Keywords + +derivative, inverse function +-/ + +universes u v w +open_locale classical topology big_operators filter ennreal +open filter asymptotics set + +variables {𝕜 : Type u} [nontrivially_normed_field 𝕜] +variables {F : Type v} [normed_add_comm_group F] [normed_space 𝕜 F] +variables {E : Type w} [normed_add_comm_group E] [normed_space 𝕜 E] + +variables {f f₀ f₁ g : 𝕜 → F} +variables {f' f₀' f₁' g' : F} +variables {x : 𝕜} +variables {s t : set 𝕜} +variables {L L₁ L₂ : filter 𝕜} + +theorem has_strict_deriv_at.has_strict_fderiv_at_equiv {f : 𝕜 → 𝕜} {f' x : 𝕜} + (hf : has_strict_deriv_at f f' x) (hf' : f' ≠ 0) : + has_strict_fderiv_at f + (continuous_linear_equiv.units_equiv_aut 𝕜 (units.mk0 f' hf') : 𝕜 →L[𝕜] 𝕜) x := +hf + +theorem has_deriv_at.has_fderiv_at_equiv {f : 𝕜 → 𝕜} {f' x : 𝕜} (hf : has_deriv_at f f' x) + (hf' : f' ≠ 0) : + has_fderiv_at f (continuous_linear_equiv.units_equiv_aut 𝕜 (units.mk0 f' hf') : 𝕜 →L[𝕜] 𝕜) x := +hf + +/-- If `f (g y) = y` for `y` in some neighborhood of `a`, `g` is continuous at `a`, and `f` has an +invertible derivative `f'` at `g a` in the strict sense, then `g` has the derivative `f'⁻¹` at `a` +in the strict sense. + +This is one of the easy parts of the inverse function theorem: it assumes that we already have an +inverse function. -/ +theorem has_strict_deriv_at.of_local_left_inverse {f g : 𝕜 → 𝕜} {f' a : 𝕜} + (hg : continuous_at g a) (hf : has_strict_deriv_at f f' (g a)) (hf' : f' ≠ 0) + (hfg : ∀ᶠ y in 𝓝 a, f (g y) = y) : + has_strict_deriv_at g f'⁻¹ a := +(hf.has_strict_fderiv_at_equiv hf').of_local_left_inverse hg hfg + +/-- If `f` is a local homeomorphism defined on a neighbourhood of `f.symm a`, and `f` has a +nonzero derivative `f'` at `f.symm a` in the strict sense, then `f.symm` has the derivative `f'⁻¹` +at `a` in the strict sense. + +This is one of the easy parts of the inverse function theorem: it assumes that we already have +an inverse function. -/ +lemma local_homeomorph.has_strict_deriv_at_symm (f : local_homeomorph 𝕜 𝕜) {a f' : 𝕜} + (ha : a ∈ f.target) (hf' : f' ≠ 0) (htff' : has_strict_deriv_at f f' (f.symm a)) : + has_strict_deriv_at f.symm f'⁻¹ a := +htff'.of_local_left_inverse (f.symm.continuous_at ha) hf' (f.eventually_right_inverse ha) + +/-- If `f (g y) = y` for `y` in some neighborhood of `a`, `g` is continuous at `a`, and `f` has an +invertible derivative `f'` at `g a`, then `g` has the derivative `f'⁻¹` at `a`. + +This is one of the easy parts of the inverse function theorem: it assumes that we already have +an inverse function. -/ +theorem has_deriv_at.of_local_left_inverse {f g : 𝕜 → 𝕜} {f' a : 𝕜} + (hg : continuous_at g a) (hf : has_deriv_at f f' (g a)) (hf' : f' ≠ 0) + (hfg : ∀ᶠ y in 𝓝 a, f (g y) = y) : + has_deriv_at g f'⁻¹ a := +(hf.has_fderiv_at_equiv hf').of_local_left_inverse hg hfg + +/-- If `f` is a local homeomorphism defined on a neighbourhood of `f.symm a`, and `f` has an +nonzero derivative `f'` at `f.symm a`, then `f.symm` has the derivative `f'⁻¹` at `a`. + +This is one of the easy parts of the inverse function theorem: it assumes that we already have +an inverse function. -/ +lemma local_homeomorph.has_deriv_at_symm (f : local_homeomorph 𝕜 𝕜) {a f' : 𝕜} + (ha : a ∈ f.target) (hf' : f' ≠ 0) (htff' : has_deriv_at f f' (f.symm a)) : + has_deriv_at f.symm f'⁻¹ a := +htff'.of_local_left_inverse (f.symm.continuous_at ha) hf' (f.eventually_right_inverse ha) + +lemma has_deriv_at.eventually_ne (h : has_deriv_at f f' x) (hf' : f' ≠ 0) : + ∀ᶠ z in 𝓝[≠] x, f z ≠ f x := +(has_deriv_at_iff_has_fderiv_at.1 h).eventually_ne + ⟨‖f'‖⁻¹, λ z, by field_simp [norm_smul, mt norm_eq_zero.1 hf']⟩ + +lemma has_deriv_at.tendsto_punctured_nhds (h : has_deriv_at f f' x) (hf' : f' ≠ 0) : + tendsto f (𝓝[≠] x) (𝓝[≠] (f x)) := +tendsto_nhds_within_of_tendsto_nhds_of_eventually_within _ + h.continuous_at.continuous_within_at (h.eventually_ne hf') + +theorem not_differentiable_within_at_of_local_left_inverse_has_deriv_within_at_zero + {f g : 𝕜 → 𝕜} {a : 𝕜} {s t : set 𝕜} (ha : a ∈ s) (hsu : unique_diff_within_at 𝕜 s a) + (hf : has_deriv_within_at f 0 t (g a)) (hst : maps_to g s t) (hfg : f ∘ g =ᶠ[𝓝[s] a] id) : + ¬differentiable_within_at 𝕜 g s a := +begin + intro hg, + have := (hf.comp a hg.has_deriv_within_at hst).congr_of_eventually_eq_of_mem hfg.symm ha, + simpa using hsu.eq_deriv _ this (has_deriv_within_at_id _ _) +end + +theorem not_differentiable_at_of_local_left_inverse_has_deriv_at_zero + {f g : 𝕜 → 𝕜} {a : 𝕜} (hf : has_deriv_at f 0 (g a)) (hfg : f ∘ g =ᶠ[𝓝 a] id) : + ¬differentiable_at 𝕜 g a := +begin + intro hg, + have := (hf.comp a hg.has_deriv_at).congr_of_eventually_eq hfg.symm, + simpa using this.unique (has_deriv_at_id a) +end diff --git a/src/analysis/calculus/deriv/linear.lean b/src/analysis/calculus/deriv/linear.lean new file mode 100644 index 0000000000000..4b6ab842f9c5d --- /dev/null +++ b/src/analysis/calculus/deriv/linear.lean @@ -0,0 +1,83 @@ +/- +Copyright (c) 2019 Gabriel Ebner. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Gabriel Ebner, Yury Kudryashov +-/ +import analysis.calculus.deriv.basic +import analysis.calculus.fderiv.linear + +/-! +# Derivatives of continuous linear maps from the base field + +In this file we prove that `f : 𝕜 →L[𝕜] E` (or `f : 𝕜 →ₗ[𝕜] E`) has derivative `f 1`. + +For a more detailed overview of one-dimensional derivatives in mathlib, see the module docstring of +`analysis/calculus/deriv/basic`. + +## Keywords + +derivative, linear map +-/ + +universes u v w + +open_locale topology filter +open filter asymptotics set + +variables {𝕜 : Type u} [nontrivially_normed_field 𝕜] +variables {F : Type v} [normed_add_comm_group F] [normed_space 𝕜 F] +variables {E : Type w} [normed_add_comm_group E] [normed_space 𝕜 E] +variables {x : 𝕜} +variables {s : set 𝕜} +variables {L : filter 𝕜} + +section continuous_linear_map +/-! ### Derivative of continuous linear maps -/ +variables (e : 𝕜 →L[𝕜] F) + +protected lemma continuous_linear_map.has_deriv_at_filter : has_deriv_at_filter e (e 1) x L := +e.has_fderiv_at_filter.has_deriv_at_filter + +protected lemma continuous_linear_map.has_strict_deriv_at : has_strict_deriv_at e (e 1) x := +e.has_strict_fderiv_at.has_strict_deriv_at + +protected lemma continuous_linear_map.has_deriv_at : has_deriv_at e (e 1) x := +e.has_deriv_at_filter + +protected lemma continuous_linear_map.has_deriv_within_at : has_deriv_within_at e (e 1) s x := +e.has_deriv_at_filter + +@[simp] protected lemma continuous_linear_map.deriv : deriv e x = e 1 := +e.has_deriv_at.deriv + +protected lemma continuous_linear_map.deriv_within (hxs : unique_diff_within_at 𝕜 s x) : + deriv_within e s x = e 1 := +e.has_deriv_within_at.deriv_within hxs + +end continuous_linear_map + +section linear_map +/-! ### Derivative of bundled linear maps -/ +variables (e : 𝕜 →ₗ[𝕜] F) + +protected lemma linear_map.has_deriv_at_filter : has_deriv_at_filter e (e 1) x L := +e.to_continuous_linear_map₁.has_deriv_at_filter + +protected lemma linear_map.has_strict_deriv_at : has_strict_deriv_at e (e 1) x := +e.to_continuous_linear_map₁.has_strict_deriv_at + +protected lemma linear_map.has_deriv_at : has_deriv_at e (e 1) x := +e.has_deriv_at_filter + +protected lemma linear_map.has_deriv_within_at : has_deriv_within_at e (e 1) s x := +e.has_deriv_at_filter + +@[simp] protected lemma linear_map.deriv : deriv e x = e 1 := +e.has_deriv_at.deriv + +protected lemma linear_map.deriv_within (hxs : unique_diff_within_at 𝕜 s x) : + deriv_within e s x = e 1 := +e.has_deriv_within_at.deriv_within hxs + +end linear_map + diff --git a/src/analysis/calculus/deriv/mul.lean b/src/analysis/calculus/deriv/mul.lean new file mode 100644 index 0000000000000..c99cc0e0b637b --- /dev/null +++ b/src/analysis/calculus/deriv/mul.lean @@ -0,0 +1,398 @@ +/- +Copyright (c) 2019 Gabriel Ebner. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Gabriel Ebner, Anatole Dedecker, Yury Kudryashov +-/ +import analysis.calculus.deriv.basic +import analysis.calculus.fderiv.mul +import analysis.calculus.fderiv.add + +/-! +# Derivative of `f x * g x` + +In this file we prove formulas for `(f x * g x)'` and `(f x • g x)'`. + +For a more detailed overview of one-dimensional derivatives in mathlib, see the module docstring of +`analysis/calculus/deriv/basic`. + +## Keywords + +derivative, multiplication +-/ + +universes u v w +noncomputable theory +open_locale classical topology big_operators filter ennreal +open filter asymptotics set +open continuous_linear_map (smul_right smul_right_one_eq_iff) + +variables {𝕜 : Type u} [nontrivially_normed_field 𝕜] +variables {F : Type v} [normed_add_comm_group F] [normed_space 𝕜 F] +variables {E : Type w} [normed_add_comm_group E] [normed_space 𝕜 E] + +variables {f f₀ f₁ g : 𝕜 → F} +variables {f' f₀' f₁' g' : F} +variables {x : 𝕜} +variables {s t : set 𝕜} +variables {L L₁ L₂ : filter 𝕜} + +section smul + +/-! ### Derivative of the multiplication of a scalar function and a vector function -/ + +variables {𝕜' : Type*} [nontrivially_normed_field 𝕜'] [normed_algebra 𝕜 𝕜'] + [normed_space 𝕜' F] [is_scalar_tower 𝕜 𝕜' F] {c : 𝕜 → 𝕜'} {c' : 𝕜'} + +theorem has_deriv_within_at.smul + (hc : has_deriv_within_at c c' s x) (hf : has_deriv_within_at f f' s x) : + has_deriv_within_at (λ y, c y • f y) (c x • f' + c' • f x) s x := +by simpa using (has_fderiv_within_at.smul hc hf).has_deriv_within_at + +theorem has_deriv_at.smul + (hc : has_deriv_at c c' x) (hf : has_deriv_at f f' x) : + has_deriv_at (λ y, c y • f y) (c x • f' + c' • f x) x := +begin + rw [← has_deriv_within_at_univ] at *, + exact hc.smul hf +end + +theorem has_strict_deriv_at.smul + (hc : has_strict_deriv_at c c' x) (hf : has_strict_deriv_at f f' x) : + has_strict_deriv_at (λ y, c y • f y) (c x • f' + c' • f x) x := +by simpa using (hc.smul hf).has_strict_deriv_at + +lemma deriv_within_smul (hxs : unique_diff_within_at 𝕜 s x) + (hc : differentiable_within_at 𝕜 c s x) (hf : differentiable_within_at 𝕜 f s x) : + deriv_within (λ y, c y • f y) s x = c x • deriv_within f s x + (deriv_within c s x) • f x := +(hc.has_deriv_within_at.smul hf.has_deriv_within_at).deriv_within hxs + +lemma deriv_smul (hc : differentiable_at 𝕜 c x) (hf : differentiable_at 𝕜 f x) : + deriv (λ y, c y • f y) x = c x • deriv f x + (deriv c x) • f x := +(hc.has_deriv_at.smul hf.has_deriv_at).deriv + +theorem has_strict_deriv_at.smul_const + (hc : has_strict_deriv_at c c' x) (f : F) : + has_strict_deriv_at (λ y, c y • f) (c' • f) x := +begin + have := hc.smul (has_strict_deriv_at_const x f), + rwa [smul_zero, zero_add] at this, +end + +theorem has_deriv_within_at.smul_const + (hc : has_deriv_within_at c c' s x) (f : F) : + has_deriv_within_at (λ y, c y • f) (c' • f) s x := +begin + have := hc.smul (has_deriv_within_at_const x s f), + rwa [smul_zero, zero_add] at this +end + +theorem has_deriv_at.smul_const + (hc : has_deriv_at c c' x) (f : F) : + has_deriv_at (λ y, c y • f) (c' • f) x := +begin + rw [← has_deriv_within_at_univ] at *, + exact hc.smul_const f +end + +lemma deriv_within_smul_const (hxs : unique_diff_within_at 𝕜 s x) + (hc : differentiable_within_at 𝕜 c s x) (f : F) : + deriv_within (λ y, c y • f) s x = (deriv_within c s x) • f := +(hc.has_deriv_within_at.smul_const f).deriv_within hxs + +lemma deriv_smul_const (hc : differentiable_at 𝕜 c x) (f : F) : + deriv (λ y, c y • f) x = (deriv c x) • f := +(hc.has_deriv_at.smul_const f).deriv + +end smul + +section const_smul + +variables {R : Type*} [semiring R] [module R F] [smul_comm_class 𝕜 R F] + [has_continuous_const_smul R F] + +theorem has_strict_deriv_at.const_smul + (c : R) (hf : has_strict_deriv_at f f' x) : + has_strict_deriv_at (λ y, c • f y) (c • f') x := +by simpa using (hf.const_smul c).has_strict_deriv_at + +theorem has_deriv_at_filter.const_smul + (c : R) (hf : has_deriv_at_filter f f' x L) : + has_deriv_at_filter (λ y, c • f y) (c • f') x L := +by simpa using (hf.const_smul c).has_deriv_at_filter + +theorem has_deriv_within_at.const_smul + (c : R) (hf : has_deriv_within_at f f' s x) : + has_deriv_within_at (λ y, c • f y) (c • f') s x := +hf.const_smul c + +theorem has_deriv_at.const_smul (c : R) (hf : has_deriv_at f f' x) : + has_deriv_at (λ y, c • f y) (c • f') x := +hf.const_smul c + +lemma deriv_within_const_smul (hxs : unique_diff_within_at 𝕜 s x) + (c : R) (hf : differentiable_within_at 𝕜 f s x) : + deriv_within (λ y, c • f y) s x = c • deriv_within f s x := +(hf.has_deriv_within_at.const_smul c).deriv_within hxs + +lemma deriv_const_smul (c : R) (hf : differentiable_at 𝕜 f x) : + deriv (λ y, c • f y) x = c • deriv f x := +(hf.has_deriv_at.const_smul c).deriv + +end const_smul + +section mul +/-! ### Derivative of the multiplication of two functions -/ +variables {𝕜' 𝔸 : Type*} [normed_field 𝕜'] [normed_ring 𝔸] [normed_algebra 𝕜 𝕜'] + [normed_algebra 𝕜 𝔸] {c d : 𝕜 → 𝔸} {c' d' : 𝔸} {u v : 𝕜 → 𝕜'} + +theorem has_deriv_within_at.mul + (hc : has_deriv_within_at c c' s x) (hd : has_deriv_within_at d d' s x) : + has_deriv_within_at (λ y, c y * d y) (c' * d x + c x * d') s x := +begin + have := (has_fderiv_within_at.mul' hc hd).has_deriv_within_at, + rwa [continuous_linear_map.add_apply, continuous_linear_map.smul_apply, + continuous_linear_map.smul_right_apply, continuous_linear_map.smul_right_apply, + continuous_linear_map.smul_right_apply, continuous_linear_map.one_apply, + one_smul, one_smul, add_comm] at this, +end + +theorem has_deriv_at.mul (hc : has_deriv_at c c' x) (hd : has_deriv_at d d' x) : + has_deriv_at (λ y, c y * d y) (c' * d x + c x * d') x := +begin + rw [← has_deriv_within_at_univ] at *, + exact hc.mul hd +end + +theorem has_strict_deriv_at.mul + (hc : has_strict_deriv_at c c' x) (hd : has_strict_deriv_at d d' x) : + has_strict_deriv_at (λ y, c y * d y) (c' * d x + c x * d') x := +begin + have := (has_strict_fderiv_at.mul' hc hd).has_strict_deriv_at, + rwa [continuous_linear_map.add_apply, continuous_linear_map.smul_apply, + continuous_linear_map.smul_right_apply, continuous_linear_map.smul_right_apply, + continuous_linear_map.smul_right_apply, continuous_linear_map.one_apply, + one_smul, one_smul, add_comm] at this, +end + +lemma deriv_within_mul (hxs : unique_diff_within_at 𝕜 s x) + (hc : differentiable_within_at 𝕜 c s x) (hd : differentiable_within_at 𝕜 d s x) : + deriv_within (λ y, c y * d y) s x = deriv_within c s x * d x + c x * deriv_within d s x := +(hc.has_deriv_within_at.mul hd.has_deriv_within_at).deriv_within hxs + +@[simp] lemma deriv_mul (hc : differentiable_at 𝕜 c x) (hd : differentiable_at 𝕜 d x) : + deriv (λ y, c y * d y) x = deriv c x * d x + c x * deriv d x := +(hc.has_deriv_at.mul hd.has_deriv_at).deriv + +theorem has_deriv_within_at.mul_const (hc : has_deriv_within_at c c' s x) (d : 𝔸) : + has_deriv_within_at (λ y, c y * d) (c' * d) s x := +begin + convert hc.mul (has_deriv_within_at_const x s d), + rw [mul_zero, add_zero] +end + +theorem has_deriv_at.mul_const (hc : has_deriv_at c c' x) (d : 𝔸) : + has_deriv_at (λ y, c y * d) (c' * d) x := +begin + rw [← has_deriv_within_at_univ] at *, + exact hc.mul_const d +end + +theorem has_deriv_at_mul_const (c : 𝕜) : has_deriv_at (λ x, x * c) c x := +by simpa only [one_mul] using (has_deriv_at_id' x).mul_const c + +theorem has_strict_deriv_at.mul_const (hc : has_strict_deriv_at c c' x) (d : 𝔸) : + has_strict_deriv_at (λ y, c y * d) (c' * d) x := +begin + convert hc.mul (has_strict_deriv_at_const x d), + rw [mul_zero, add_zero] +end + +lemma deriv_within_mul_const (hxs : unique_diff_within_at 𝕜 s x) + (hc : differentiable_within_at 𝕜 c s x) (d : 𝔸) : + deriv_within (λ y, c y * d) s x = deriv_within c s x * d := +(hc.has_deriv_within_at.mul_const d).deriv_within hxs + +lemma deriv_mul_const (hc : differentiable_at 𝕜 c x) (d : 𝔸) : + deriv (λ y, c y * d) x = deriv c x * d := +(hc.has_deriv_at.mul_const d).deriv + +lemma deriv_mul_const_field (v : 𝕜') : + deriv (λ y, u y * v) x = deriv u x * v := +begin + by_cases hu : differentiable_at 𝕜 u x, + { exact deriv_mul_const hu v }, + { rw [deriv_zero_of_not_differentiable_at hu, zero_mul], + rcases eq_or_ne v 0 with rfl|hd, + { simp only [mul_zero, deriv_const] }, + { refine deriv_zero_of_not_differentiable_at (mt (λ H, _) hu), + simpa only [mul_inv_cancel_right₀ hd] using H.mul_const v⁻¹ } } +end + +@[simp] lemma deriv_mul_const_field' (v : 𝕜') : deriv (λ x, u x * v) = λ x, deriv u x * v := +funext $ λ _, deriv_mul_const_field v + +theorem has_deriv_within_at.const_mul (c : 𝔸) (hd : has_deriv_within_at d d' s x) : + has_deriv_within_at (λ y, c * d y) (c * d') s x := +begin + convert (has_deriv_within_at_const x s c).mul hd, + rw [zero_mul, zero_add] +end + +theorem has_deriv_at.const_mul (c : 𝔸) (hd : has_deriv_at d d' x) : + has_deriv_at (λ y, c * d y) (c * d') x := +begin + rw [← has_deriv_within_at_univ] at *, + exact hd.const_mul c +end + +theorem has_strict_deriv_at.const_mul (c : 𝔸) (hd : has_strict_deriv_at d d' x) : + has_strict_deriv_at (λ y, c * d y) (c * d') x := +begin + convert (has_strict_deriv_at_const _ _).mul hd, + rw [zero_mul, zero_add] +end + +lemma deriv_within_const_mul (hxs : unique_diff_within_at 𝕜 s x) + (c : 𝔸) (hd : differentiable_within_at 𝕜 d s x) : + deriv_within (λ y, c * d y) s x = c * deriv_within d s x := +(hd.has_deriv_within_at.const_mul c).deriv_within hxs + +lemma deriv_const_mul (c : 𝔸) (hd : differentiable_at 𝕜 d x) : + deriv (λ y, c * d y) x = c * deriv d x := +(hd.has_deriv_at.const_mul c).deriv + +lemma deriv_const_mul_field (u : 𝕜') : deriv (λ y, u * v y) x = u * deriv v x := +by simp only [mul_comm u, deriv_mul_const_field] + +@[simp] lemma deriv_const_mul_field' (u : 𝕜') : deriv (λ x, u * v x) = λ x, u * deriv v x := +funext (λ x, deriv_const_mul_field u) + +end mul + +section div + + +variables {𝕜' : Type*} [nontrivially_normed_field 𝕜'] [normed_algebra 𝕜 𝕜'] + {c d : 𝕜 → 𝕜'} {c' d' : 𝕜'} + +lemma has_deriv_at.div_const (hc : has_deriv_at c c' x) (d : 𝕜') : + has_deriv_at (λ x, c x / d) (c' / d) x := +by simpa only [div_eq_mul_inv] using hc.mul_const d⁻¹ + +lemma has_deriv_within_at.div_const (hc : has_deriv_within_at c c' s x) (d : 𝕜') : + has_deriv_within_at (λ x, c x / d) (c' / d) s x := +by simpa only [div_eq_mul_inv] using hc.mul_const d⁻¹ + +lemma has_strict_deriv_at.div_const (hc : has_strict_deriv_at c c' x) (d : 𝕜') : + has_strict_deriv_at (λ x, c x / d) (c' / d) x := +by simpa only [div_eq_mul_inv] using hc.mul_const d⁻¹ + +lemma differentiable_within_at.div_const (hc : differentiable_within_at 𝕜 c s x) (d : 𝕜') : + differentiable_within_at 𝕜 (λx, c x / d) s x := +(hc.has_deriv_within_at.div_const _).differentiable_within_at + +@[simp] lemma differentiable_at.div_const (hc : differentiable_at 𝕜 c x) (d : 𝕜') : + differentiable_at 𝕜 (λ x, c x / d) x := +(hc.has_deriv_at.div_const _).differentiable_at + +lemma differentiable_on.div_const (hc : differentiable_on 𝕜 c s) (d : 𝕜') : + differentiable_on 𝕜 (λx, c x / d) s := +λ x hx, (hc x hx).div_const d + +@[simp] lemma differentiable.div_const (hc : differentiable 𝕜 c) (d : 𝕜') : + differentiable 𝕜 (λx, c x / d) := +λ x, (hc x).div_const d + +lemma deriv_within_div_const (hc : differentiable_within_at 𝕜 c s x) (d : 𝕜') + (hxs : unique_diff_within_at 𝕜 s x) : + deriv_within (λx, c x / d) s x = (deriv_within c s x) / d := +by simp [div_eq_inv_mul, deriv_within_const_mul, hc, hxs] + +@[simp] lemma deriv_div_const (d : 𝕜') : + deriv (λx, c x / d) x = (deriv c x) / d := +by simp only [div_eq_mul_inv, deriv_mul_const_field] + +end div + +section clm_comp_apply +/-! ### Derivative of the pointwise composition/application of continuous linear maps -/ + +open continuous_linear_map + +variables {G : Type*} [normed_add_comm_group G] [normed_space 𝕜 G] {c : 𝕜 → F →L[𝕜] G} + {c' : F →L[𝕜] G} {d : 𝕜 → E →L[𝕜] F} {d' : E →L[𝕜] F} {u : 𝕜 → F} {u' : F} + +lemma has_strict_deriv_at.clm_comp (hc : has_strict_deriv_at c c' x) + (hd : has_strict_deriv_at d d' x) : + has_strict_deriv_at (λ y, (c y).comp (d y)) (c'.comp (d x) + (c x).comp d') x := +begin + have := (hc.has_strict_fderiv_at.clm_comp hd.has_strict_fderiv_at).has_strict_deriv_at, + rwa [add_apply, comp_apply, comp_apply, smul_right_apply, smul_right_apply, one_apply, one_smul, + one_smul, add_comm] at this, +end + +lemma has_deriv_within_at.clm_comp (hc : has_deriv_within_at c c' s x) + (hd : has_deriv_within_at d d' s x) : + has_deriv_within_at (λ y, (c y).comp (d y)) (c'.comp (d x) + (c x).comp d') s x := +begin + have := (hc.has_fderiv_within_at.clm_comp hd.has_fderiv_within_at).has_deriv_within_at, + rwa [add_apply, comp_apply, comp_apply, smul_right_apply, smul_right_apply, one_apply, one_smul, + one_smul, add_comm] at this, +end + +lemma has_deriv_at.clm_comp (hc : has_deriv_at c c' x) (hd : has_deriv_at d d' x) : + has_deriv_at (λ y, (c y).comp (d y)) + (c'.comp (d x) + (c x).comp d') x := +begin + rw [← has_deriv_within_at_univ] at *, + exact hc.clm_comp hd +end + +lemma deriv_within_clm_comp (hc : differentiable_within_at 𝕜 c s x) + (hd : differentiable_within_at 𝕜 d s x) (hxs : unique_diff_within_at 𝕜 s x): + deriv_within (λ y, (c y).comp (d y)) s x = + ((deriv_within c s x).comp (d x) + (c x).comp (deriv_within d s x)) := +(hc.has_deriv_within_at.clm_comp hd.has_deriv_within_at).deriv_within hxs + +lemma deriv_clm_comp (hc : differentiable_at 𝕜 c x) (hd : differentiable_at 𝕜 d x) : + deriv (λ y, (c y).comp (d y)) x = + ((deriv c x).comp (d x) + (c x).comp (deriv d x)) := +(hc.has_deriv_at.clm_comp hd.has_deriv_at).deriv + +lemma has_strict_deriv_at.clm_apply (hc : has_strict_deriv_at c c' x) + (hu : has_strict_deriv_at u u' x) : + has_strict_deriv_at (λ y, (c y) (u y)) (c' (u x) + c x u') x := +begin + have := (hc.has_strict_fderiv_at.clm_apply hu.has_strict_fderiv_at).has_strict_deriv_at, + rwa [add_apply, comp_apply, flip_apply, smul_right_apply, smul_right_apply, one_apply, one_smul, + one_smul, add_comm] at this, +end + +lemma has_deriv_within_at.clm_apply (hc : has_deriv_within_at c c' s x) + (hu : has_deriv_within_at u u' s x) : + has_deriv_within_at (λ y, (c y) (u y)) (c' (u x) + c x u') s x := +begin + have := (hc.has_fderiv_within_at.clm_apply hu.has_fderiv_within_at).has_deriv_within_at, + rwa [add_apply, comp_apply, flip_apply, smul_right_apply, smul_right_apply, one_apply, one_smul, + one_smul, add_comm] at this, +end + +lemma has_deriv_at.clm_apply (hc : has_deriv_at c c' x) (hu : has_deriv_at u u' x) : + has_deriv_at (λ y, (c y) (u y)) (c' (u x) + c x u') x := +begin + have := (hc.has_fderiv_at.clm_apply hu.has_fderiv_at).has_deriv_at, + rwa [add_apply, comp_apply, flip_apply, smul_right_apply, smul_right_apply, one_apply, one_smul, + one_smul, add_comm] at this, +end + +lemma deriv_within_clm_apply (hxs : unique_diff_within_at 𝕜 s x) + (hc : differentiable_within_at 𝕜 c s x) (hu : differentiable_within_at 𝕜 u s x) : + deriv_within (λ y, (c y) (u y)) s x = (deriv_within c s x (u x) + c x (deriv_within u s x)) := +(hc.has_deriv_within_at.clm_apply hu.has_deriv_within_at).deriv_within hxs + +lemma deriv_clm_apply (hc : differentiable_at 𝕜 c x) (hu : differentiable_at 𝕜 u x) : + deriv (λ y, (c y) (u y)) x = (deriv c x (u x) + c x (deriv u x)) := +(hc.has_deriv_at.clm_apply hu.has_deriv_at).deriv + +end clm_comp_apply + diff --git a/src/analysis/calculus/deriv/polynomial.lean b/src/analysis/calculus/deriv/polynomial.lean new file mode 100644 index 0000000000000..089e604c39f31 --- /dev/null +++ b/src/analysis/calculus/deriv/polynomial.lean @@ -0,0 +1,157 @@ +/- +Copyright (c) 2019 Sébastien Gouëzel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sébastien Gouëzel, Eric Wieser +-/ +import analysis.calculus.deriv.pow +import analysis.calculus.deriv.add +import data.polynomial.algebra_map +import data.polynomial.derivative + +/-! +# Derivatives of polynomials + +In this file we prove that derivatives of polynomials in the analysis sense agree with their +derivatives in the algebraic sense. + +For a more detailed overview of one-dimensional derivatives in mathlib, see the module docstring of +`analysis/calculus/deriv/basic`. + +## TODO + +* Add results about multivariable polynomials. +* Generalize some (most?) results to an algebra over the base field. + +## Keywords + +derivative, polynomial +-/ + +universes u v w +open_locale classical topology big_operators filter ennreal polynomial +open filter asymptotics set +open continuous_linear_map (smul_right smul_right_one_eq_iff) + + +variables {𝕜 : Type u} [nontrivially_normed_field 𝕜] +variables {F : Type v} [normed_add_comm_group F] [normed_space 𝕜 F] +variables {E : Type w} [normed_add_comm_group E] [normed_space 𝕜 E] + +variables {f f₀ f₁ g : 𝕜 → F} +variables {f' f₀' f₁' g' : F} +variables {x : 𝕜} +variables {s t : set 𝕜} +variables {L L₁ L₂ : filter 𝕜} + +namespace polynomial +/-! ### Derivative of a polynomial -/ + +variables {R : Type*} [comm_semiring R] [algebra R 𝕜] +variables (p : 𝕜[X]) (q : R[X]) + +/-- The derivative (in the analysis sense) of a polynomial `p` is given by `p.derivative`. -/ +protected lemma has_strict_deriv_at (x : 𝕜) : + has_strict_deriv_at (λx, p.eval x) (p.derivative.eval x) x := +begin + induction p using polynomial.induction_on', + case h_add : p q hp hq { simpa using hp.add hq }, + case h_monomial : n a { simpa [mul_assoc] using (has_strict_deriv_at_pow n x).const_mul a } +end + +protected lemma has_strict_deriv_at_aeval (x : 𝕜) : + has_strict_deriv_at (λx, aeval x q) (aeval x q.derivative) x := +by simpa only [aeval_def, eval₂_eq_eval_map, derivative_map] + using (q.map (algebra_map R 𝕜)).has_strict_deriv_at x + +/-- The derivative (in the analysis sense) of a polynomial `p` is given by `p.derivative`. -/ +protected lemma has_deriv_at (x : 𝕜) : has_deriv_at (λx, p.eval x) (p.derivative.eval x) x := +(p.has_strict_deriv_at x).has_deriv_at + +protected lemma has_deriv_at_aeval (x : 𝕜) : + has_deriv_at (λx, aeval x q) (aeval x q.derivative) x := +(q.has_strict_deriv_at_aeval x).has_deriv_at + +protected theorem has_deriv_within_at (x : 𝕜) (s : set 𝕜) : + has_deriv_within_at (λx, p.eval x) (p.derivative.eval x) s x := +(p.has_deriv_at x).has_deriv_within_at + +protected theorem has_deriv_within_at_aeval (x : 𝕜) (s : set 𝕜) : + has_deriv_within_at (λx, aeval x q) (aeval x q.derivative) s x := +(q.has_deriv_at_aeval x).has_deriv_within_at + +protected lemma differentiable_at : differentiable_at 𝕜 (λx, p.eval x) x := +(p.has_deriv_at x).differentiable_at + +protected lemma differentiable_at_aeval : differentiable_at 𝕜 (λx, aeval x q) x := +(q.has_deriv_at_aeval x).differentiable_at + +protected lemma differentiable_within_at : differentiable_within_at 𝕜 (λx, p.eval x) s x := +p.differentiable_at.differentiable_within_at + +protected lemma differentiable_within_at_aeval : differentiable_within_at 𝕜 (λx, aeval x q) s x := +q.differentiable_at_aeval.differentiable_within_at + +protected lemma differentiable : differentiable 𝕜 (λx, p.eval x) := +λx, p.differentiable_at + +protected lemma differentiable_aeval : differentiable 𝕜 (λ x : 𝕜, aeval x q) := +λx, q.differentiable_at_aeval + +protected lemma differentiable_on : differentiable_on 𝕜 (λx, p.eval x) s := +p.differentiable.differentiable_on + +protected lemma differentiable_on_aeval : differentiable_on 𝕜 (λx, aeval x q) s := +q.differentiable_aeval.differentiable_on + +@[simp] protected lemma deriv : deriv (λx, p.eval x) x = p.derivative.eval x := +(p.has_deriv_at x).deriv + +@[simp] protected lemma deriv_aeval : deriv (λx, aeval x q) x = aeval x q.derivative := +(q.has_deriv_at_aeval x).deriv + +protected lemma deriv_within (hxs : unique_diff_within_at 𝕜 s x) : + deriv_within (λx, p.eval x) s x = p.derivative.eval x := +begin + rw differentiable_at.deriv_within p.differentiable_at hxs, + exact p.deriv +end + +protected lemma deriv_within_aeval (hxs : unique_diff_within_at 𝕜 s x) : + deriv_within (λx, aeval x q) s x = aeval x q.derivative := +by simpa only [aeval_def, eval₂_eq_eval_map, derivative_map] + using (q.map (algebra_map R 𝕜)).deriv_within hxs + +protected lemma has_fderiv_at (x : 𝕜) : + has_fderiv_at (λx, p.eval x) (smul_right (1 : 𝕜 →L[𝕜] 𝕜) (p.derivative.eval x)) x := +p.has_deriv_at x + +protected lemma has_fderiv_at_aeval (x : 𝕜) : + has_fderiv_at (λx, aeval x q) (smul_right (1 : 𝕜 →L[𝕜] 𝕜) (aeval x q.derivative)) x := +q.has_deriv_at_aeval x + +protected lemma has_fderiv_within_at (x : 𝕜) : + has_fderiv_within_at (λx, p.eval x) (smul_right (1 : 𝕜 →L[𝕜] 𝕜) (p.derivative.eval x)) s x := +(p.has_fderiv_at x).has_fderiv_within_at + +protected lemma has_fderiv_within_at_aeval (x : 𝕜) : + has_fderiv_within_at (λx, aeval x q) (smul_right (1 : 𝕜 →L[𝕜] 𝕜) (aeval x q.derivative)) s x := +(q.has_fderiv_at_aeval x).has_fderiv_within_at + +@[simp] protected lemma fderiv : + fderiv 𝕜 (λx, p.eval x) x = smul_right (1 : 𝕜 →L[𝕜] 𝕜) (p.derivative.eval x) := +(p.has_fderiv_at x).fderiv + +@[simp] protected lemma fderiv_aeval : + fderiv 𝕜 (λx, aeval x q) x = smul_right (1 : 𝕜 →L[𝕜] 𝕜) (aeval x q.derivative) := +(q.has_fderiv_at_aeval x).fderiv + +protected lemma fderiv_within (hxs : unique_diff_within_at 𝕜 s x) : + fderiv_within 𝕜 (λx, p.eval x) s x = smul_right (1 : 𝕜 →L[𝕜] 𝕜) (p.derivative.eval x) := +(p.has_fderiv_within_at x).fderiv_within hxs + +protected lemma fderiv_within_aeval (hxs : unique_diff_within_at 𝕜 s x) : + fderiv_within 𝕜 (λx, aeval x q) s x = smul_right (1 : 𝕜 →L[𝕜] 𝕜) (aeval x q.derivative) := +(q.has_fderiv_within_at_aeval x).fderiv_within hxs + +end polynomial + diff --git a/src/analysis/calculus/deriv/pow.lean b/src/analysis/calculus/deriv/pow.lean new file mode 100644 index 0000000000000..d1c0b5796657c --- /dev/null +++ b/src/analysis/calculus/deriv/pow.lean @@ -0,0 +1,91 @@ +/- +Copyright (c) 2019 Sébastien Gouëzel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sébastien Gouëzel +-/ +import analysis.calculus.deriv.mul +import analysis.calculus.deriv.comp + +/-! +# Derivative of `(f x) ^ n`, `n : ℕ` + +In this file we prove that `(x ^ n)' = n * x ^ (n - 1)`, where `n` is a natural number. + +For a more detailed overview of one-dimensional derivatives in mathlib, see the module docstring of +`analysis/calculus/deriv/basic`. + +## Keywords + +derivative, power +-/ + +universes u v w +open_locale classical topology big_operators filter ennreal +open filter asymptotics set + +variables {𝕜 : Type u} [nontrivially_normed_field 𝕜] +variables {F : Type v} [normed_add_comm_group F] [normed_space 𝕜 F] +variables {E : Type w} [normed_add_comm_group E] [normed_space 𝕜 E] + +variables {f f₀ f₁ g : 𝕜 → F} +variables {f' f₀' f₁' g' : F} +variables {x : 𝕜} +variables {s t : set 𝕜} +variables {L L₁ L₂ : filter 𝕜} + +/-! ### Derivative of `x ↦ x^n` for `n : ℕ` -/ +variables {c : 𝕜 → 𝕜} {c' : 𝕜} +variable (n : ℕ) + +lemma has_strict_deriv_at_pow : ∀ (n : ℕ) (x : 𝕜), + has_strict_deriv_at (λx, x^n) ((n : 𝕜) * x^(n-1)) x +| 0 x := by simp [has_strict_deriv_at_const] +| 1 x := by simpa using has_strict_deriv_at_id x +| (n + 1 + 1) x := by simpa [pow_succ', add_mul, mul_assoc] + using (has_strict_deriv_at_pow (n + 1) x).mul (has_strict_deriv_at_id x) + +lemma has_deriv_at_pow (n : ℕ) (x : 𝕜) : has_deriv_at (λx, x^n) ((n : 𝕜) * x^(n-1)) x := +(has_strict_deriv_at_pow n x).has_deriv_at + +theorem has_deriv_within_at_pow (n : ℕ) (x : 𝕜) (s : set 𝕜) : + has_deriv_within_at (λx, x^n) ((n : 𝕜) * x^(n-1)) s x := +(has_deriv_at_pow n x).has_deriv_within_at + +lemma differentiable_at_pow : differentiable_at 𝕜 (λx, x^n) x := +(has_deriv_at_pow n x).differentiable_at + +lemma differentiable_within_at_pow : differentiable_within_at 𝕜 (λx, x^n) s x := +(differentiable_at_pow n).differentiable_within_at + +lemma differentiable_pow : differentiable 𝕜 (λx:𝕜, x^n) := +λ x, differentiable_at_pow n + +lemma differentiable_on_pow : differentiable_on 𝕜 (λx, x^n) s := +(differentiable_pow n).differentiable_on + +lemma deriv_pow : deriv (λ x, x^n) x = (n : 𝕜) * x^(n-1) := +(has_deriv_at_pow n x).deriv + +@[simp] lemma deriv_pow' : deriv (λ x, x^n) = λ x, (n : 𝕜) * x^(n-1) := +funext $ λ x, deriv_pow n + +lemma deriv_within_pow (hxs : unique_diff_within_at 𝕜 s x) : + deriv_within (λx, x^n) s x = (n : 𝕜) * x^(n-1) := +(has_deriv_within_at_pow n x s).deriv_within hxs + +lemma has_deriv_within_at.pow (hc : has_deriv_within_at c c' s x) : + has_deriv_within_at (λ y, (c y)^n) ((n : 𝕜) * (c x)^(n-1) * c') s x := +(has_deriv_at_pow n (c x)).comp_has_deriv_within_at x hc + +lemma has_deriv_at.pow (hc : has_deriv_at c c' x) : + has_deriv_at (λ y, (c y)^n) ((n : 𝕜) * (c x)^(n-1) * c') x := +by { rw ← has_deriv_within_at_univ at *, exact hc.pow n } + +lemma deriv_within_pow' (hc : differentiable_within_at 𝕜 c s x) + (hxs : unique_diff_within_at 𝕜 s x) : + deriv_within (λx, (c x)^n) s x = (n : 𝕜) * (c x)^(n-1) * (deriv_within c s x) := +(hc.has_deriv_within_at.pow n).deriv_within hxs + +@[simp] lemma deriv_pow'' (hc : differentiable_at 𝕜 c x) : + deriv (λx, (c x)^n) x = (n : 𝕜) * (c x)^(n-1) * (deriv c x) := +(hc.has_deriv_at.pow n).deriv diff --git a/src/analysis/calculus/deriv/prod.lean b/src/analysis/calculus/deriv/prod.lean new file mode 100644 index 0000000000000..2ab31ee57d62c --- /dev/null +++ b/src/analysis/calculus/deriv/prod.lean @@ -0,0 +1,98 @@ +/- +Copyright (c) 2019 Gabriel Ebner. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Gabriel Ebner, Yury Kudryashov +-/ +import analysis.calculus.deriv.basic +import analysis.calculus.fderiv.prod + +/-! +# Derivatives of functions taking values in product types + +In this file we prove lemmas about derivatives of functions `f : 𝕜 → E × F` and of functions +`f : 𝕜 → (Π i, E i)`. + +For a more detailed overview of one-dimensional derivatives in mathlib, see the module docstring of +`analysis/calculus/deriv/basic`. + +## Keywords + +derivative +-/ + +universes u v w +open_locale classical topology big_operators filter +open filter asymptotics set + +variables {𝕜 : Type u} [nontrivially_normed_field 𝕜] +variables {F : Type v} [normed_add_comm_group F] [normed_space 𝕜 F] +variables {E : Type w} [normed_add_comm_group E] [normed_space 𝕜 E] + +variables {f f₀ f₁ g : 𝕜 → F} +variables {f' f₀' f₁' g' : F} +variables {x : 𝕜} +variables {s t : set 𝕜} +variables {L L₁ L₂ : filter 𝕜} + +section cartesian_product +/-! ### Derivative of the cartesian product of two functions -/ + +variables {G : Type w} [normed_add_comm_group G] [normed_space 𝕜 G] +variables {f₂ : 𝕜 → G} {f₂' : G} + +lemma has_deriv_at_filter.prod + (hf₁ : has_deriv_at_filter f₁ f₁' x L) (hf₂ : has_deriv_at_filter f₂ f₂' x L) : + has_deriv_at_filter (λ x, (f₁ x, f₂ x)) (f₁', f₂') x L := +hf₁.prod hf₂ + +lemma has_deriv_within_at.prod + (hf₁ : has_deriv_within_at f₁ f₁' s x) (hf₂ : has_deriv_within_at f₂ f₂' s x) : + has_deriv_within_at (λ x, (f₁ x, f₂ x)) (f₁', f₂') s x := +hf₁.prod hf₂ + +lemma has_deriv_at.prod (hf₁ : has_deriv_at f₁ f₁' x) (hf₂ : has_deriv_at f₂ f₂' x) : + has_deriv_at (λ x, (f₁ x, f₂ x)) (f₁', f₂') x := +hf₁.prod hf₂ + +lemma has_strict_deriv_at.prod (hf₁ : has_strict_deriv_at f₁ f₁' x) + (hf₂ : has_strict_deriv_at f₂ f₂' x) : + has_strict_deriv_at (λ x, (f₁ x, f₂ x)) (f₁', f₂') x := +hf₁.prod hf₂ + +end cartesian_product + +section pi + +/-! ### Derivatives of functions `f : 𝕜 → Π i, E i` -/ + +variables {ι : Type*} [fintype ι] {E' : ι → Type*} [Π i, normed_add_comm_group (E' i)] + [Π i, normed_space 𝕜 (E' i)] {φ : 𝕜 → Π i, E' i} {φ' : Π i, E' i} + +@[simp] lemma has_strict_deriv_at_pi : + has_strict_deriv_at φ φ' x ↔ ∀ i, has_strict_deriv_at (λ x, φ x i) (φ' i) x := +has_strict_fderiv_at_pi' + +@[simp] lemma has_deriv_at_filter_pi : + has_deriv_at_filter φ φ' x L ↔ + ∀ i, has_deriv_at_filter (λ x, φ x i) (φ' i) x L := +has_fderiv_at_filter_pi' + +lemma has_deriv_at_pi : + has_deriv_at φ φ' x ↔ ∀ i, has_deriv_at (λ x, φ x i) (φ' i) x:= +has_deriv_at_filter_pi + +lemma has_deriv_within_at_pi : + has_deriv_within_at φ φ' s x ↔ ∀ i, has_deriv_within_at (λ x, φ x i) (φ' i) s x:= +has_deriv_at_filter_pi + +lemma deriv_within_pi (h : ∀ i, differentiable_within_at 𝕜 (λ x, φ x i) s x) + (hs : unique_diff_within_at 𝕜 s x) : + deriv_within φ s x = λ i, deriv_within (λ x, φ x i) s x := +(has_deriv_within_at_pi.2 (λ i, (h i).has_deriv_within_at)).deriv_within hs + +lemma deriv_pi (h : ∀ i, differentiable_at 𝕜 (λ x, φ x i) x) : + deriv φ x = λ i, deriv (λ x, φ x i) x := +(has_deriv_at_pi.2 (λ i, (h i).has_deriv_at)).deriv + +end pi + diff --git a/src/analysis/calculus/deriv/slope.lean b/src/analysis/calculus/deriv/slope.lean new file mode 100644 index 0000000000000..7f883efca6727 --- /dev/null +++ b/src/analysis/calculus/deriv/slope.lean @@ -0,0 +1,178 @@ +/- +Copyright (c) 2019 Yury Kudryashov. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yury Kudryashov +-/ +import analysis.calculus.deriv.basic +import linear_algebra.affine_space.slope + +/-! +# Derivative as the limit of the slope + +In this file we relate the derivative of a function with its definition from a standard +undergraduate course as the limit of the slope `(f y - f x) / (y - x)` as `y` tends to `𝓝[≠] x`. +Since we are talking about functions taking values in a normed space instead of the base field, we +use `slope f x y = (y - x)⁻¹ • (f y - f x)` instead of division. + +We also prove some estimates on the upper/lower limits of the slope in terms of the derivative. + +For a more detailed overview of one-dimensional derivatives in mathlib, see the module docstring of +`analysis/calculus/deriv/basic`. + +## Keywords + +derivative, slope +-/ + +universes u v w +noncomputable theory +open_locale classical topology big_operators filter ennreal +open filter asymptotics set +open continuous_linear_map (smul_right smul_right_one_eq_iff) + +section normed_field + +variables {𝕜 : Type u} [nontrivially_normed_field 𝕜] +variables {F : Type v} [normed_add_comm_group F] [normed_space 𝕜 F] +variables {E : Type w} [normed_add_comm_group E] [normed_space 𝕜 E] + +variables {f f₀ f₁ g : 𝕜 → F} +variables {f' f₀' f₁' g' : F} +variables {x : 𝕜} +variables {s t : set 𝕜} +variables {L L₁ L₂ : filter 𝕜} + +/-- If the domain has dimension one, then Fréchet derivative is equivalent to the classical +definition with a limit. In this version we have to take the limit along the subset `-{x}`, +because for `y=x` the slope equals zero due to the convention `0⁻¹=0`. -/ +lemma has_deriv_at_filter_iff_tendsto_slope {x : 𝕜} {L : filter 𝕜} : + has_deriv_at_filter f f' x L ↔ tendsto (slope f x) (L ⊓ 𝓟 {x}ᶜ) (𝓝 f') := +begin + conv_lhs { simp only [has_deriv_at_filter_iff_tendsto, (norm_inv _).symm, + (norm_smul _ _).symm, tendsto_zero_iff_norm_tendsto_zero.symm] }, + conv_rhs { rw [← nhds_translation_sub f', tendsto_comap_iff] }, + refine (tendsto_inf_principal_nhds_iff_of_forall_eq $ by simp).symm.trans (tendsto_congr' _), + refine (eventually_principal.2 $ λ z hz, _).filter_mono inf_le_right, + simp only [(∘)], + rw [smul_sub, ← mul_smul, inv_mul_cancel (sub_ne_zero.2 hz), one_smul, slope_def_module] +end + +lemma has_deriv_within_at_iff_tendsto_slope : + has_deriv_within_at f f' s x ↔ tendsto (slope f x) (𝓝[s \ {x}] x) (𝓝 f') := +begin + simp only [has_deriv_within_at, nhds_within, diff_eq, inf_assoc.symm, inf_principal.symm], + exact has_deriv_at_filter_iff_tendsto_slope +end + +lemma has_deriv_within_at_iff_tendsto_slope' (hs : x ∉ s) : + has_deriv_within_at f f' s x ↔ tendsto (slope f x) (𝓝[s] x) (𝓝 f') := +begin + convert ← has_deriv_within_at_iff_tendsto_slope, + exact diff_singleton_eq_self hs +end + +lemma has_deriv_at_iff_tendsto_slope : + has_deriv_at f f' x ↔ tendsto (slope f x) (𝓝[≠] x) (𝓝 f') := +has_deriv_at_filter_iff_tendsto_slope + +end normed_field + +/-! ### Upper estimates on liminf and limsup -/ + +section real + +variables {f : ℝ → ℝ} {f' : ℝ} {s : set ℝ} {x : ℝ} {r : ℝ} + +lemma has_deriv_within_at.limsup_slope_le (hf : has_deriv_within_at f f' s x) (hr : f' < r) : + ∀ᶠ z in 𝓝[s \ {x}] x, slope f x z < r := +has_deriv_within_at_iff_tendsto_slope.1 hf (is_open.mem_nhds is_open_Iio hr) + +lemma has_deriv_within_at.limsup_slope_le' (hf : has_deriv_within_at f f' s x) + (hs : x ∉ s) (hr : f' < r) : + ∀ᶠ z in 𝓝[s] x, slope f x z < r := +(has_deriv_within_at_iff_tendsto_slope' hs).1 hf (is_open.mem_nhds is_open_Iio hr) + +lemma has_deriv_within_at.liminf_right_slope_le + (hf : has_deriv_within_at f f' (Ici x) x) (hr : f' < r) : + ∃ᶠ z in 𝓝[>] x, slope f x z < r := +(hf.Ioi_of_Ici.limsup_slope_le' (lt_irrefl x) hr).frequently + +end real + +section real_space + +open metric + +variables {E : Type u} [normed_add_comm_group E] [normed_space ℝ E] {f : ℝ → E} {f' : E} {s : set ℝ} + {x r : ℝ} + +/-- If `f` has derivative `f'` within `s` at `x`, then for any `r > ‖f'‖` the ratio +`‖f z - f x‖ / ‖z - x‖` is less than `r` in some neighborhood of `x` within `s`. +In other words, the limit superior of this ratio as `z` tends to `x` along `s` +is less than or equal to `‖f'‖`. -/ +lemma has_deriv_within_at.limsup_norm_slope_le + (hf : has_deriv_within_at f f' s x) (hr : ‖f'‖ < r) : + ∀ᶠ z in 𝓝[s] x, ‖z - x‖⁻¹ * ‖f z - f x‖ < r := +begin + have hr₀ : 0 < r, from lt_of_le_of_lt (norm_nonneg f') hr, + have A : ∀ᶠ z in 𝓝[s \ {x}] x, ‖(z - x)⁻¹ • (f z - f x)‖ ∈ Iio r, + from (has_deriv_within_at_iff_tendsto_slope.1 hf).norm (is_open.mem_nhds is_open_Iio hr), + have B : ∀ᶠ z in 𝓝[{x}] x, ‖(z - x)⁻¹ • (f z - f x)‖ ∈ Iio r, + from mem_of_superset self_mem_nhds_within + (singleton_subset_iff.2 $ by simp [hr₀]), + have C := mem_sup.2 ⟨A, B⟩, + rw [← nhds_within_union, diff_union_self, nhds_within_union, mem_sup] at C, + filter_upwards [C.1], + simp only [norm_smul, mem_Iio, norm_inv], + exact λ _, id +end + +/-- If `f` has derivative `f'` within `s` at `x`, then for any `r > ‖f'‖` the ratio +`(‖f z‖ - ‖f x‖) / ‖z - x‖` is less than `r` in some neighborhood of `x` within `s`. +In other words, the limit superior of this ratio as `z` tends to `x` along `s` +is less than or equal to `‖f'‖`. + +This lemma is a weaker version of `has_deriv_within_at.limsup_norm_slope_le` +where `‖f z‖ - ‖f x‖` is replaced by `‖f z - f x‖`. -/ +lemma has_deriv_within_at.limsup_slope_norm_le + (hf : has_deriv_within_at f f' s x) (hr : ‖f'‖ < r) : + ∀ᶠ z in 𝓝[s] x, ‖z - x‖⁻¹ * (‖f z‖ - ‖f x‖) < r := +begin + apply (hf.limsup_norm_slope_le hr).mono, + assume z hz, + refine lt_of_le_of_lt (mul_le_mul_of_nonneg_left (norm_sub_norm_le _ _) _) hz, + exact inv_nonneg.2 (norm_nonneg _) +end + +/-- If `f` has derivative `f'` within `(x, +∞)` at `x`, then for any `r > ‖f'‖` the ratio +`‖f z - f x‖ / ‖z - x‖` is frequently less than `r` as `z → x+0`. +In other words, the limit inferior of this ratio as `z` tends to `x+0` +is less than or equal to `‖f'‖`. See also `has_deriv_within_at.limsup_norm_slope_le` +for a stronger version using limit superior and any set `s`. -/ +lemma has_deriv_within_at.liminf_right_norm_slope_le + (hf : has_deriv_within_at f f' (Ici x) x) (hr : ‖f'‖ < r) : + ∃ᶠ z in 𝓝[>] x, ‖z - x‖⁻¹ * ‖f z - f x‖ < r := +(hf.Ioi_of_Ici.limsup_norm_slope_le hr).frequently + +/-- If `f` has derivative `f'` within `(x, +∞)` at `x`, then for any `r > ‖f'‖` the ratio +`(‖f z‖ - ‖f x‖) / (z - x)` is frequently less than `r` as `z → x+0`. +In other words, the limit inferior of this ratio as `z` tends to `x+0` +is less than or equal to `‖f'‖`. + +See also + +* `has_deriv_within_at.limsup_norm_slope_le` for a stronger version using + limit superior and any set `s`; +* `has_deriv_within_at.liminf_right_norm_slope_le` for a stronger version using + `‖f z - f x‖` instead of `‖f z‖ - ‖f x‖`. -/ +lemma has_deriv_within_at.liminf_right_slope_norm_le + (hf : has_deriv_within_at f f' (Ici x) x) (hr : ‖f'‖ < r) : + ∃ᶠ z in 𝓝[>] x, (z - x)⁻¹ * (‖f z‖ - ‖f x‖) < r := +begin + have := (hf.Ioi_of_Ici.limsup_slope_norm_le hr).frequently, + refine this.mp (eventually.mono self_mem_nhds_within _), + assume z hxz hz, + rwa [real.norm_eq_abs, abs_of_pos (sub_pos_of_lt hxz)] at hz +end + +end real_space diff --git a/src/analysis/calculus/deriv/star.lean b/src/analysis/calculus/deriv/star.lean new file mode 100644 index 0000000000000..6f9ad983dd1fd --- /dev/null +++ b/src/analysis/calculus/deriv/star.lean @@ -0,0 +1,66 @@ +/- +Copyright (c) 2023 Eric Wieser. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Eric Wieser +-/ +import analysis.calculus.deriv.basic +import analysis.calculus.fderiv.star + +/-! +# Star operations on derivatives + +This file contains the usual formulas (and existence assertions) for the derivative of the star +operation. Note that these only apply when the field that the derivative is respect to has a trivial +star operation; which as should be expected rules out `𝕜 = ℂ`. +-/ + +universes u v w +noncomputable theory +open_locale classical topology big_operators filter ennreal +open filter asymptotics set +open continuous_linear_map (smul_right smul_right_one_eq_iff) + + +variables {𝕜 : Type u} [nontrivially_normed_field 𝕜] +variables {F : Type v} [normed_add_comm_group F] [normed_space 𝕜 F] +variables {E : Type w} [normed_add_comm_group E] [normed_space 𝕜 E] + +variables {f f₀ f₁ g : 𝕜 → F} +variables {f' f₀' f₁' g' : F} +variables {x : 𝕜} +variables {s t : set 𝕜} +variables {L L₁ L₂ : filter 𝕜} + +section star +/-! ### Derivative of `x ↦ star x` -/ + +variables [star_ring 𝕜] [has_trivial_star 𝕜] [star_add_monoid F] [has_continuous_star F] +variable [star_module 𝕜 F] + +protected theorem has_deriv_at_filter.star (h : has_deriv_at_filter f f' x L) : + has_deriv_at_filter (λ x, star (f x)) (star f') x L := +by simpa using h.star.has_deriv_at_filter + +protected theorem has_deriv_within_at.star (h : has_deriv_within_at f f' s x) : + has_deriv_within_at (λ x, star (f x)) (star f') s x := +h.star + +protected theorem has_deriv_at.star (h : has_deriv_at f f' x) : + has_deriv_at (λ x, star (f x)) (star f') x := +h.star + +protected theorem has_strict_deriv_at.star (h : has_strict_deriv_at f f' x) : + has_strict_deriv_at (λ x, star (f x)) (star f') x := +by simpa using h.star.has_strict_deriv_at + +protected lemma deriv_within.star (hxs : unique_diff_within_at 𝕜 s x) : + deriv_within (λ y, star (f y)) s x = star (deriv_within f s x) := +fun_like.congr_fun (fderiv_within_star hxs) _ + +protected lemma deriv.star : deriv (λ y, star (f y)) x = star (deriv f x) := +fun_like.congr_fun fderiv_star _ + +@[simp] protected lemma deriv.star' : deriv (λ y, star (f y)) = (λ x, star (deriv f x)) := +funext $ λ x, deriv.star + +end star diff --git a/src/analysis/calculus/deriv/support.lean b/src/analysis/calculus/deriv/support.lean new file mode 100644 index 0000000000000..cbf46e0ebc3fb --- /dev/null +++ b/src/analysis/calculus/deriv/support.lean @@ -0,0 +1,44 @@ +/- +Copyright (c) 2022 Floris van Doorn. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Floris van Doorn +-/ +import analysis.calculus.deriv.basic + +/-! +# Support of the derivative of a function + +In this file we prove that the (topological) support of a function includes the support of its +derivative. As a corollary, we show that the derivative of a function with compact support has +compact support. + +## Keywords + +derivative, support +-/ + +universes u v + +variables {𝕜 : Type u} [nontrivially_normed_field 𝕜] +variables {E : Type v} [normed_add_comm_group E] [normed_space 𝕜 E] +variables {f : 𝕜 → E} + +/-! ### Support of derivatives -/ + +section support + +open function + +lemma support_deriv_subset : support (deriv f) ⊆ tsupport f := +begin + intros x, + rw [← not_imp_not], + intro h2x, + rw [not_mem_tsupport_iff_eventually_eq] at h2x, + exact nmem_support.mpr (h2x.deriv_eq.trans (deriv_const x 0)) +end + +lemma has_compact_support.deriv (hf : has_compact_support f) : has_compact_support (deriv f) := +hf.mono' support_deriv_subset + +end support diff --git a/src/analysis/calculus/deriv/zpow.lean b/src/analysis/calculus/deriv/zpow.lean new file mode 100644 index 0000000000000..867c779eb3b23 --- /dev/null +++ b/src/analysis/calculus/deriv/zpow.lean @@ -0,0 +1,149 @@ +/- +Copyright (c) 2020 Sébastien Gouëzel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sébastien Gouëzel, Yury Kudryashov +-/ +import analysis.calculus.deriv.pow +import analysis.calculus.deriv.inv + +/-! +# Derivatives of `x ^ m`, `m : ℤ` + +In this file we prove theorems about (iterated) derivatives of `x ^ m`, `m : ℤ`. + +For a more detailed overview of one-dimensional derivatives in mathlib, see the module docstring of +`analysis/calculus/deriv/basic`. + +## Keywords + +derivative, power +-/ + +universes u v w +open_locale classical topology big_operators filter +open filter asymptotics set + +variables {𝕜 : Type u} [nontrivially_normed_field 𝕜] +variables {E : Type v} [normed_add_comm_group E] [normed_space 𝕜 E] + +variables {x : 𝕜} +variables {s : set 𝕜} +variables {m : ℤ} + +/-! ### Derivative of `x ↦ x^m` for `m : ℤ` -/ + +lemma has_strict_deriv_at_zpow (m : ℤ) (x : 𝕜) (h : x ≠ 0 ∨ 0 ≤ m) : + has_strict_deriv_at (λx, x^m) ((m : 𝕜) * x^(m-1)) x := +begin + have : ∀ m : ℤ, 0 < m → has_strict_deriv_at (λx, x^m) ((m:𝕜) * x^(m-1)) x, + { assume m hm, + lift m to ℕ using (le_of_lt hm), + simp only [zpow_coe_nat, int.cast_coe_nat], + convert has_strict_deriv_at_pow _ _ using 2, + rw [← int.coe_nat_one, ← int.coe_nat_sub, zpow_coe_nat], + norm_cast at hm, + exact nat.succ_le_of_lt hm }, + rcases lt_trichotomy m 0 with hm|hm|hm, + { have hx : x ≠ 0, from h.resolve_right hm.not_le, + have := (has_strict_deriv_at_inv _).scomp _ (this (-m) (neg_pos.2 hm)); + [skip, exact zpow_ne_zero_of_ne_zero hx _], + simp only [(∘), zpow_neg, one_div, inv_inv, smul_eq_mul] at this, + convert this using 1, + rw [sq, mul_inv, inv_inv, int.cast_neg, neg_mul, neg_mul_neg, + ← zpow_add₀ hx, mul_assoc, ← zpow_add₀ hx], congr, abel }, + { simp only [hm, zpow_zero, int.cast_zero, zero_mul, has_strict_deriv_at_const] }, + { exact this m hm } +end + +lemma has_deriv_at_zpow (m : ℤ) (x : 𝕜) (h : x ≠ 0 ∨ 0 ≤ m) : + has_deriv_at (λx, x^m) ((m : 𝕜) * x^(m-1)) x := +(has_strict_deriv_at_zpow m x h).has_deriv_at + +theorem has_deriv_within_at_zpow (m : ℤ) (x : 𝕜) (h : x ≠ 0 ∨ 0 ≤ m) (s : set 𝕜) : + has_deriv_within_at (λx, x^m) ((m : 𝕜) * x^(m-1)) s x := +(has_deriv_at_zpow m x h).has_deriv_within_at + +lemma differentiable_at_zpow : differentiable_at 𝕜 (λx, x^m) x ↔ x ≠ 0 ∨ 0 ≤ m := +⟨λ H, normed_field.continuous_at_zpow.1 H.continuous_at, + λ H, (has_deriv_at_zpow m x H).differentiable_at⟩ + +lemma differentiable_within_at_zpow (m : ℤ) (x : 𝕜) (h : x ≠ 0 ∨ 0 ≤ m) : + differentiable_within_at 𝕜 (λx, x^m) s x := +(differentiable_at_zpow.mpr h).differentiable_within_at + +lemma differentiable_on_zpow (m : ℤ) (s : set 𝕜) (h : (0 : 𝕜) ∉ s ∨ 0 ≤ m) : + differentiable_on 𝕜 (λx, x^m) s := +λ x hxs, differentiable_within_at_zpow m x $ h.imp_left $ ne_of_mem_of_not_mem hxs + +lemma deriv_zpow (m : ℤ) (x : 𝕜) : deriv (λ x, x ^ m) x = m * x ^ (m - 1) := +begin + by_cases H : x ≠ 0 ∨ 0 ≤ m, + { exact (has_deriv_at_zpow m x H).deriv }, + { rw deriv_zero_of_not_differentiable_at (mt differentiable_at_zpow.1 H), + push_neg at H, rcases H with ⟨rfl, hm⟩, + rw [zero_zpow _ ((sub_one_lt _).trans hm).ne, mul_zero] } +end + +@[simp] lemma deriv_zpow' (m : ℤ) : deriv (λ x : 𝕜, x ^ m) = λ x, m * x ^ (m - 1) := +funext $ deriv_zpow m + +lemma deriv_within_zpow (hxs : unique_diff_within_at 𝕜 s x) (h : x ≠ 0 ∨ 0 ≤ m) : + deriv_within (λx, x^m) s x = (m : 𝕜) * x^(m-1) := +(has_deriv_within_at_zpow m x h s).deriv_within hxs + +@[simp] lemma iter_deriv_zpow' (m : ℤ) (k : ℕ) : + deriv^[k] (λ x : 𝕜, x ^ m) = λ x, (∏ i in finset.range k, (m - i)) * x ^ (m - k) := +begin + induction k with k ihk, + { simp only [one_mul, int.coe_nat_zero, id, sub_zero, finset.prod_range_zero, + function.iterate_zero] }, + { simp only [function.iterate_succ_apply', ihk, deriv_const_mul_field', deriv_zpow', + finset.prod_range_succ, int.coe_nat_succ, ← sub_sub, int.cast_sub, int.cast_coe_nat, + mul_assoc], } +end + +lemma iter_deriv_zpow (m : ℤ) (x : 𝕜) (k : ℕ) : + deriv^[k] (λ y, y ^ m) x = (∏ i in finset.range k, (m - i)) * x ^ (m - k) := +congr_fun (iter_deriv_zpow' m k) x + +lemma iter_deriv_pow (n : ℕ) (x : 𝕜) (k : ℕ) : + deriv^[k] (λx:𝕜, x^n) x = (∏ i in finset.range k, (n - i)) * x^(n-k) := +begin + simp only [← zpow_coe_nat, iter_deriv_zpow, int.cast_coe_nat], + cases le_or_lt k n with hkn hnk, + { rw int.coe_nat_sub hkn }, + { have : ∏ i in finset.range k, (n - i : 𝕜) = 0, + from finset.prod_eq_zero (finset.mem_range.2 hnk) (sub_self _), + simp only [this, zero_mul] } +end + +@[simp] lemma iter_deriv_pow' (n k : ℕ) : + deriv^[k] (λ x : 𝕜, x ^ n) = λ x, (∏ i in finset.range k, (n - i)) * x ^ (n - k) := +funext $ λ x, iter_deriv_pow n x k + +lemma iter_deriv_inv (k : ℕ) (x : 𝕜) : + deriv^[k] has_inv.inv x = (∏ i in finset.range k, (-1 - i)) * x ^ (-1 - k : ℤ) := +by simpa only [zpow_neg_one, int.cast_neg, int.cast_one] using iter_deriv_zpow (-1) x k + +@[simp] lemma iter_deriv_inv' (k : ℕ) : + deriv^[k] has_inv.inv = λ x : 𝕜, (∏ i in finset.range k, (-1 - i)) * x ^ (-1 - k : ℤ) := +funext (iter_deriv_inv k) + +variables {f : E → 𝕜} {t : set E} {a : E} + +lemma differentiable_within_at.zpow (hf : differentiable_within_at 𝕜 f t a) (h : f a ≠ 0 ∨ 0 ≤ m) : + differentiable_within_at 𝕜 (λ x, f x ^ m) t a := +(differentiable_at_zpow.2 h).comp_differentiable_within_at a hf + +lemma differentiable_at.zpow (hf : differentiable_at 𝕜 f a) (h : f a ≠ 0 ∨ 0 ≤ m) : + differentiable_at 𝕜 (λ x, f x ^ m) a := +(differentiable_at_zpow.2 h).comp a hf + +lemma differentiable_on.zpow (hf : differentiable_on 𝕜 f t) (h : (∀ x ∈ t, f x ≠ 0) ∨ 0 ≤ m) : + differentiable_on 𝕜 (λ x, f x ^ m) t := +λ x hx, (hf x hx).zpow $ h.imp_left (λ h, h x hx) + +lemma differentiable.zpow (hf : differentiable 𝕜 f) (h : (∀ x, f x ≠ 0) ∨ 0 ≤ m) : + differentiable 𝕜 (λ x, f x ^ m) := +λ x, (hf x).zpow $ h.imp_left (λ h, h x) + diff --git a/src/analysis/calculus/diff_cont_on_cl.lean b/src/analysis/calculus/diff_cont_on_cl.lean index 2ca22622a955b..4150d47e822da 100644 --- a/src/analysis/calculus/diff_cont_on_cl.lean +++ b/src/analysis/calculus/diff_cont_on_cl.lean @@ -3,7 +3,7 @@ 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 analysis.calculus.deriv +import analysis.calculus.deriv.inv /-! # Functions differentiable on a domain and continuous on its closure diff --git a/src/analysis/calculus/dslope.lean b/src/analysis/calculus/dslope.lean index 48b5af6941003..810aa968ddc1d 100644 --- a/src/analysis/calculus/dslope.lean +++ b/src/analysis/calculus/dslope.lean @@ -3,7 +3,8 @@ Copyright (c) 2022 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ -import analysis.calculus.deriv +import analysis.calculus.deriv.slope +import analysis.calculus.deriv.inv /-! # Slope of a differentiable function diff --git a/src/analysis/calculus/fderiv_analytic.lean b/src/analysis/calculus/fderiv_analytic.lean index bddec64dbb2bc..9b2775354f2fe 100644 --- a/src/analysis/calculus/fderiv_analytic.lean +++ b/src/analysis/calculus/fderiv_analytic.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import analysis.analytic.basic -import analysis.calculus.deriv +import analysis.calculus.deriv.basic import analysis.calculus.cont_diff_def /-! diff --git a/src/analysis/calculus/fderiv_measurable.lean b/src/analysis/calculus/fderiv_measurable.lean index ed414fa04a82b..562df4994997c 100644 --- a/src/analysis/calculus/fderiv_measurable.lean +++ b/src/analysis/calculus/fderiv_measurable.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel, Yury Kudryashov -/ -import analysis.calculus.deriv +import analysis.calculus.deriv.basic import measure_theory.constructions.borel_space.continuous_linear_map import measure_theory.function.strongly_measurable.basic diff --git a/src/analysis/calculus/iterated_deriv.lean b/src/analysis/calculus/iterated_deriv.lean index 58d836ff92a25..634a49a1ec745 100644 --- a/src/analysis/calculus/iterated_deriv.lean +++ b/src/analysis/calculus/iterated_deriv.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ -import analysis.calculus.deriv +import analysis.calculus.deriv.comp import analysis.calculus.cont_diff_def /-! diff --git a/src/analysis/calculus/lhopital.lean b/src/analysis/calculus/lhopital.lean index 200fd62607420..5c2b6d617e122 100644 --- a/src/analysis/calculus/lhopital.lean +++ b/src/analysis/calculus/lhopital.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Anatole Dedecker -/ import analysis.calculus.mean_value +import analysis.calculus.deriv.inv /-! # L'Hôpital's rule for 0/0 indeterminate forms diff --git a/src/analysis/calculus/local_extr.lean b/src/analysis/calculus/local_extr.lean index ef494a5da1bcb..f0fb3bcf9859d 100644 --- a/src/analysis/calculus/local_extr.lean +++ b/src/analysis/calculus/local_extr.lean @@ -3,7 +3,7 @@ Copyright (c) 2019 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ -import analysis.calculus.deriv +import analysis.calculus.deriv.polynomial import topology.algebra.order.extend_from import topology.algebra.polynomial diff --git a/src/analysis/calculus/mean_value.lean b/src/analysis/calculus/mean_value.lean index 959f372245ec1..5b38db38f1e77 100644 --- a/src/analysis/calculus/mean_value.lean +++ b/src/analysis/calculus/mean_value.lean @@ -3,6 +3,7 @@ Copyright (c) 2019 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel, Yury Kudryashov -/ +import analysis.calculus.deriv.slope import analysis.calculus.local_extr import analysis.convex.slope import analysis.convex.normed diff --git a/src/analysis/calculus/monotone.lean b/src/analysis/calculus/monotone.lean index c94163209a778..d20ad597d8a53 100644 --- a/src/analysis/calculus/monotone.lean +++ b/src/analysis/calculus/monotone.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ -import analysis.calculus.deriv +import analysis.calculus.deriv.slope import measure_theory.covering.one_dim import order.monotone.extension diff --git a/src/analysis/complex/real_deriv.lean b/src/analysis/complex/real_deriv.lean index f3cf1837f66c2..b88e9735e27cc 100644 --- a/src/analysis/complex/real_deriv.lean +++ b/src/analysis/complex/real_deriv.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel, Yourong Zang -/ import analysis.calculus.cont_diff +import analysis.calculus.deriv.linear import analysis.complex.conformal import analysis.calculus.conformal.normed_space diff --git a/src/analysis/convex/specific_functions/deriv.lean b/src/analysis/convex/specific_functions/deriv.lean index c7ba9d71ba84d..907264c088293 100644 --- a/src/analysis/convex/specific_functions/deriv.lean +++ b/src/analysis/convex/specific_functions/deriv.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov, Sébastien Gouëzel -/ import analysis.convex.specific_functions.basic +import analysis.calculus.deriv.zpow import analysis.special_functions.pow.deriv import analysis.special_functions.sqrt import tactic.linear_combination diff --git a/src/analysis/special_functions/log/deriv.lean b/src/analysis/special_functions/log/deriv.lean index 8db6780d53702..cc1b00d9942ed 100644 --- a/src/analysis/special_functions/log/deriv.lean +++ b/src/analysis/special_functions/log/deriv.lean @@ -3,6 +3,8 @@ 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.calculus.deriv.pow +import analysis.calculus.deriv.inv import analysis.special_functions.log.basic import analysis.special_functions.exp_deriv diff --git a/src/analysis/special_functions/pow/deriv.lean b/src/analysis/special_functions/pow/deriv.lean index e122335f31c1b..ee55bb82b775d 100644 --- a/src/analysis/special_functions/pow/deriv.lean +++ b/src/analysis/special_functions/pow/deriv.lean @@ -7,6 +7,7 @@ Authors: Chris Hughes, Abhimanyu Pallavi Sudhir, Jean Lo, Calle Sönne, Sébasti import analysis.special_functions.pow.continuity import analysis.special_functions.complex.log_deriv import analysis.calculus.extend_deriv +import analysis.calculus.deriv.prod import analysis.special_functions.log.deriv import analysis.special_functions.trigonometric.deriv diff --git a/src/geometry/manifold/instances/sphere.lean b/src/geometry/manifold/instances/sphere.lean index 785bef9a7930a..23bae7aa168e2 100644 --- a/src/geometry/manifold/instances/sphere.lean +++ b/src/geometry/manifold/instances/sphere.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Heather Macbeth. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Heather Macbeth -/ +import analysis.calculus.deriv.inv import analysis.normed_space.ball_action import analysis.special_functions.exp_deriv import analysis.inner_product_space.calculus diff --git a/src/measure_theory/integral/circle_integral.lean b/src/measure_theory/integral/circle_integral.lean index 9ae1dcff76095..80470acd02a2f 100644 --- a/src/measure_theory/integral/circle_integral.lean +++ b/src/measure_theory/integral/circle_integral.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import measure_theory.integral.interval_integral +import analysis.calculus.deriv.zpow import analysis.normed_space.pointwise import analysis.special_functions.non_integrable import analysis.analytic.basic diff --git a/src/measure_theory/integral/divergence_theorem.lean b/src/measure_theory/integral/divergence_theorem.lean index f0b9fa18f81d5..5ec6d78cadbce 100644 --- a/src/measure_theory/integral/divergence_theorem.lean +++ b/src/measure_theory/integral/divergence_theorem.lean @@ -5,7 +5,7 @@ Authors: Yury Kudryashov -/ import analysis.box_integral.divergence_theorem import analysis.box_integral.integrability -import analysis.calculus.deriv +import analysis.calculus.deriv.basic import measure_theory.constructions.prod.integral import measure_theory.integral.interval_integral diff --git a/src/measure_theory/integral/fund_thm_calculus.lean b/src/measure_theory/integral/fund_thm_calculus.lean index 43b4c4da776a4..d19811f6d32cb 100644 --- a/src/measure_theory/integral/fund_thm_calculus.lean +++ b/src/measure_theory/integral/fund_thm_calculus.lean @@ -4,6 +4,10 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury G. Kudryashov, Patrick Massot, Sébastien Gouëzel -/ import analysis.calculus.fderiv_measurable +import analysis.calculus.deriv.comp +import analysis.calculus.deriv.add +import analysis.calculus.deriv.slope +import analysis.calculus.deriv.mul import analysis.normed_space.dual import measure_theory.integral.interval_integral import measure_theory.integral.vitali_caratheodory diff --git a/src/ring_theory/polynomial/hermite/gaussian.lean b/src/ring_theory/polynomial/hermite/gaussian.lean index 0dc26ea76c07f..261e7de7865eb 100644 --- a/src/ring_theory/polynomial/hermite/gaussian.lean +++ b/src/ring_theory/polynomial/hermite/gaussian.lean @@ -3,8 +3,9 @@ Copyright (c) 2023 Luke Mantle. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Luke Mantle, Jake Levinson -/ - import ring_theory.polynomial.hermite.basic +import analysis.calculus.deriv.pow +import analysis.calculus.deriv.add import analysis.special_functions.exp import analysis.special_functions.exp_deriv diff --git a/test/simp_command.lean b/test/simp_command.lean index 6601174e96c32..83d2a5d2c877f 100644 --- a/test/simp_command.lean +++ b/test/simp_command.lean @@ -1,5 +1,6 @@ import tactic.simp_command import analysis.special_functions.trigonometric.deriv +import analysis.calculus.deriv.inv /- Turn off trace messages only if the statements are simplified to true: -/ set_option trace.silence_simp_if_true true From 917c3c072e487b3cccdbfeff17e75b40e45f66cb Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Sat, 27 May 2023 20:56:06 +0000 Subject: [PATCH 1100/1291] feat (analysis/mellin_transform): API for mellin transforms (#19110) This adds results about Mellin transforms of things like `f(a * x)` or `f (x ^ a)` in terms of the Mellin transform of `f`. --- src/analysis/mellin_transform.lean | 222 ++++++++++++++++-- .../special_functions/gamma/basic.lean | 2 +- 2 files changed, 201 insertions(+), 23 deletions(-) diff --git a/src/analysis/mellin_transform.lean b/src/analysis/mellin_transform.lean index fc32cd4d4c28f..c3aae84253868 100644 --- a/src/analysis/mellin_transform.lean +++ b/src/analysis/mellin_transform.lean @@ -6,6 +6,7 @@ Authors: David Loeffler import analysis.special_functions.improper_integrals import analysis.calculus.parametric_integral +import measure_theory.measure.haar.normed_space /-! # The Mellin transform @@ -16,6 +17,8 @@ differentiable in a suitable vertical strip. - `mellin` : the Mellin transform `∫ (t : ℝ) in Ioi 0, t ^ (s - 1) • f t`, where `s` is a complex number. +- `has_mellin`: shorthand asserting that the Mellin transform exists and has a given value + (analogous to `has_sum`). - `mellin_differentiable_at_of_is_O_rpow` : if `f` is `O(x ^ (-a))` at infinity, and `O(x ^ (-b))` at 0, then `mellin f` is holomorphic on the domain `b < re s < a`. @@ -23,22 +26,162 @@ differentiable in a suitable vertical strip. open measure_theory set filter asymptotics topological_space +namespace complex + +/- Porting note: move this to `analysis.special_functions.pow.complex` -/ +lemma cpow_mul_of_real_nonneg {x : ℝ} (hx : 0 ≤ x) (y : ℝ) (z : ℂ) : + (x : ℂ) ^ (↑y * z) = (↑(x ^ y) : ℂ) ^ z := +begin + rw [cpow_mul, of_real_cpow hx], + { rw [←of_real_log hx, ←of_real_mul, of_real_im, neg_lt_zero], exact real.pi_pos }, + { rw [←of_real_log hx, ←of_real_mul, of_real_im], exact real.pi_pos.le }, +end + +end complex + +open real complex (hiding exp log abs_of_nonneg) + open_locale topology noncomputable theory section defs -variables {E : Type*} [normed_add_comm_group E] +variables {E : Type*} [normed_add_comm_group E] [normed_space ℂ E] + +/-- Predicate on `f` and `s` asserting that the Mellin integral is well-defined. -/ +def mellin_convergent (f : ℝ → E) (s : ℂ) : Prop := +integrable_on (λ t : ℝ, (t : ℂ) ^ (s - 1) • f t) (Ioi 0) + +lemma mellin_convergent.const_smul {f : ℝ → E} {s : ℂ} (hf : mellin_convergent f s) + {𝕜 : Type*} [nontrivially_normed_field 𝕜] [normed_space 𝕜 E] [smul_comm_class ℂ 𝕜 E] (c : 𝕜) : + mellin_convergent (λ t, c • f t) s := +by simpa only [mellin_convergent, smul_comm] using hf.smul c + +lemma mellin_convergent.cpow_smul {f : ℝ → E} {s a : ℂ} : + mellin_convergent (λ t, (t : ℂ) ^ a • f t) s ↔ mellin_convergent f (s + a) := +begin + refine integrable_on_congr_fun (λ t ht, _) measurable_set_Ioi, + simp_rw [←sub_add_eq_add_sub, cpow_add _ _ (of_real_ne_zero.2 $ ne_of_gt ht), mul_smul], +end + +lemma mellin_convergent.div_const {f : ℝ → ℂ} {s : ℂ} (hf : mellin_convergent f s) (a : ℂ) : + mellin_convergent (λ t, f t / a) s := +by simpa only [mellin_convergent, smul_eq_mul, ←mul_div_assoc] using hf.div_const a + +lemma mellin_convergent.comp_mul_left {f : ℝ → E} {s : ℂ} {a : ℝ} (ha : 0 < a) : + mellin_convergent (λ t, f (a * t)) s ↔ mellin_convergent f s := +begin + have := integrable_on_Ioi_comp_mul_left_iff (λ t : ℝ, (t : ℂ) ^ (s - 1) • f t) 0 ha, + rw mul_zero at this, + have h1 : eq_on (λ t : ℝ, (↑(a * t) : ℂ) ^ (s - 1) • f (a * t)) + ((a : ℂ) ^ (s - 1) • (λ t : ℝ, (t : ℂ) ^ (s - 1) • f (a * t))) (Ioi 0), + { intros t ht, + simp only [of_real_mul, mul_cpow_of_real_nonneg ha.le (le_of_lt ht), mul_smul, pi.smul_apply] }, + have h2 : (a : ℂ) ^ (s - 1) ≠ 0, + { rw [ne.def, cpow_eq_zero_iff, not_and_distrib, of_real_eq_zero], exact or.inl ha.ne' }, + simp_rw [mellin_convergent, ←this, integrable_on_congr_fun h1 measurable_set_Ioi, integrable_on, + integrable_smul_iff h2], +end + +lemma mellin_convergent.comp_rpow {f : ℝ → E} {s : ℂ} {a : ℝ} (ha : a ≠ 0) : + mellin_convergent (λ t, f (t ^ a)) s ↔ mellin_convergent f (s / a) := +begin + simp_rw mellin_convergent, + letI u : normed_space ℝ E := normed_space.complex_to_real, -- why isn't this automatic? + conv_rhs { rw ←@integrable_on_Ioi_comp_rpow_iff' _ _ u _ a ha }, + refine integrable_on_congr_fun (λ t ht, _) measurable_set_Ioi, + dsimp only [pi.smul_apply], + rw [←complex.coe_smul (t ^ (a - 1)), ←mul_smul, ←cpow_mul_of_real_nonneg (le_of_lt ht), + of_real_cpow (le_of_lt ht), ←cpow_add _ _ (of_real_ne_zero.mpr (ne_of_gt ht)), of_real_sub, + of_real_one, mul_sub, mul_div_cancel' _ (of_real_ne_zero.mpr ha), mul_one, add_comm, + ←add_sub_assoc, sub_add_cancel], +end + +variables [complete_space E] /-- The Mellin transform of a function `f` (for a complex exponent `s`), defined as the integral of `t ^ (s - 1) • f` over `Ioi 0`. -/ -def mellin [normed_space ℂ E] [complete_space E] (f : ℝ → E) (s : ℂ) : E := +def mellin (f : ℝ → E) (s : ℂ) : E := ∫ t : ℝ in Ioi 0, (t : ℂ) ^ (s - 1) • f t -end defs +-- next few lemmas don't require convergence of the Mellin transform (they are just 0 = 0 otherwise) -open real complex (hiding exp log abs_of_nonneg) +lemma mellin_cpow_smul (f : ℝ → E) (s a : ℂ) : + mellin (λ t, (t : ℂ) ^ a • f t) s = mellin f (s + a) := +begin + refine set_integral_congr measurable_set_Ioi (λ t ht, _), + simp_rw [←sub_add_eq_add_sub, cpow_add _ _ (of_real_ne_zero.2 $ ne_of_gt ht), mul_smul], +end + +lemma mellin_const_smul (f : ℝ → E) (s : ℂ) + {𝕜 : Type*} [nontrivially_normed_field 𝕜] [normed_space 𝕜 E] [smul_comm_class ℂ 𝕜 E] (c : 𝕜) : + mellin (λ t, c • f t) s = c • mellin f s := +by simp only [mellin, smul_comm, integral_smul] + +lemma mellin_div_const (f : ℝ → ℂ) (s a : ℂ) : + mellin (λ t, f t / a) s = mellin f s / a := +by simp_rw [mellin, smul_eq_mul, ←mul_div_assoc, integral_div] + +lemma mellin_comp_rpow (f : ℝ → E) (s : ℂ) {a : ℝ} (ha : a ≠ 0) : + mellin (λ t, f (t ^ a)) s = |a|⁻¹ • mellin f (s / a) := +begin + -- note: this is also true for a = 0 (both sides are zero), but this is mathematically + -- uninteresting and rather time-consuming to check + simp_rw mellin, + conv_rhs { rw [←integral_comp_rpow_Ioi _ ha, ←integral_smul] }, + refine set_integral_congr measurable_set_Ioi (λ t ht, _), + dsimp only, + rw [←mul_smul, ←mul_assoc, inv_mul_cancel, one_mul, ←smul_assoc, real_smul], + show |a| ≠ 0, { contrapose! ha, exact abs_eq_zero.mp ha }, + rw [of_real_cpow (le_of_lt ht), ←cpow_mul_of_real_nonneg (le_of_lt ht), + ←cpow_add _ _ (of_real_ne_zero.mpr $ ne_of_gt ht), of_real_sub, of_real_one, mul_sub, + mul_div_cancel' _ (of_real_ne_zero.mpr ha), add_comm, ←add_sub_assoc, mul_one, sub_add_cancel] +end + +lemma mellin_comp_mul_left (f : ℝ → E) (s : ℂ) {a : ℝ} (ha : 0 < a) : + mellin (λ t, f (a * t)) s = (a : ℂ) ^ (-s) • mellin f s := +begin + simp_rw mellin, + have : eq_on (λ t : ℝ, (t : ℂ) ^ (s - 1) • f (a * t)) + (λ t : ℝ, (a : ℂ) ^ (1 - s) • (λ u : ℝ, (u : ℂ) ^ (s - 1) • f u) (a * t)) (Ioi 0), + { intros t ht, + dsimp only, + rw [of_real_mul, mul_cpow_of_real_nonneg ha.le (le_of_lt ht), ←mul_smul, + (by ring : 1 - s = -(s - 1)), cpow_neg, inv_mul_cancel_left₀], + rw [ne.def, cpow_eq_zero_iff, of_real_eq_zero, not_and_distrib], + exact or.inl ha.ne' }, + rw [set_integral_congr measurable_set_Ioi this, integral_smul, + integral_comp_mul_left_Ioi _ _ ha, mul_zero, ←complex.coe_smul, ←mul_smul, sub_eq_add_neg, + cpow_add _ _ (of_real_ne_zero.mpr ha.ne'), cpow_one, abs_of_pos (inv_pos.mpr ha), of_real_inv, + mul_assoc, mul_comm, inv_mul_cancel_right₀ (of_real_ne_zero.mpr ha.ne')] +end + +lemma mellin_comp_mul_right (f : ℝ → E) (s : ℂ) {a : ℝ} (ha : 0 < a) : + mellin (λ t, f (t * a)) s = (a : ℂ) ^ (-s) • mellin f s := +by simpa only [mul_comm] using mellin_comp_mul_left f s ha + +lemma mellin_comp_inv (f : ℝ → E) (s : ℂ) : mellin (λ t, f (t⁻¹)) s = mellin f (-s) := +by simp_rw [←rpow_neg_one, mellin_comp_rpow _ _ (neg_ne_zero.mpr one_ne_zero), abs_neg, abs_one, + inv_one, one_smul, of_real_neg, of_real_one, div_neg, div_one] + +/-- Predicate standing for "the Mellin transform of `f` is defined at `s` and equal to `m`". This +shortens some arguments. -/ +def has_mellin (f : ℝ → E) (s : ℂ) (m : E) : Prop := mellin_convergent f s ∧ mellin f s = m + +lemma has_mellin_add {f g : ℝ → E} {s : ℂ} + (hf : mellin_convergent f s) (hg : mellin_convergent g s) : + has_mellin (λ t, f t + g t) s (mellin f s + mellin g s) := +⟨by simpa only [mellin_convergent, smul_add] using hf.add hg, + by simpa only [mellin, smul_add] using integral_add hf hg⟩ + +lemma has_mellin_sub {f g : ℝ → E} {s : ℂ} + (hf : mellin_convergent f s) (hg : mellin_convergent g s) : + has_mellin (λ t, f t - g t) s (mellin f s - mellin g s) := +⟨by simpa only [mellin_convergent, smul_sub] using hf.sub hg, + by simpa only [mellin, smul_sub] using integral_sub hf hg⟩ + +end defs variables {E : Type*} [normed_add_comm_group E] @@ -92,7 +235,7 @@ end `b < s`, its Mellin transform converges on some right neighbourhood of `0`. -/ lemma mellin_convergent_zero_of_is_O {b : ℝ} {f : ℝ → ℝ} (hfc : ae_strongly_measurable f $ volume.restrict (Ioi 0)) - (hf : is_O (𝓝[Ioi 0] 0) f (λ t, t ^ (-b))) {s : ℝ} (hs : b < s) : + (hf : is_O (𝓝[>] 0) f (λ t, t ^ (-b))) {s : ℝ} (hs : b < s) : ∃ (c : ℝ), 0 < c ∧ integrable_on (λ t : ℝ, t ^ (s - 1) * f t) (Ioc 0 c) := begin obtain ⟨d, hd, hd'⟩ := hf.exists_pos, @@ -127,7 +270,7 @@ lemma mellin_convergent_of_is_O_scalar {a b : ℝ} {f : ℝ → ℝ} {s : ℝ} (hfc : locally_integrable_on f $ Ioi 0) (hf_top : is_O at_top f (λ t, t ^ (-a))) (hs_top : s < a) - (hf_bot : is_O (𝓝[Ioi 0] 0) f (λ t, t ^ (-b))) (hs_bot : b < s) : + (hf_bot : is_O (𝓝[>] 0) f (λ t, t ^ (-b))) (hs_bot : b < s) : integrable_on (λ t : ℝ, t ^ (s - 1) * f t) (Ioi 0) := begin obtain ⟨c1, hc1, hc1'⟩ := mellin_convergent_top_of_is_O hfc.ae_strongly_measurable hf_top hs_top, @@ -146,11 +289,11 @@ lemma mellin_convergent_of_is_O_rpow [normed_space ℂ E] {a b : ℝ} {f : ℝ → E} {s : ℂ} (hfc : locally_integrable_on f $ Ioi 0) (hf_top : is_O at_top f (λ t, t ^ (-a))) (hs_top : s.re < a) - (hf_bot : is_O (𝓝[Ioi 0] 0) f (λ t, t ^ (-b))) (hs_bot : b < s.re) : - integrable_on (λ t : ℝ, (t : ℂ) ^ (s - 1) • f t) (Ioi 0) := + (hf_bot : is_O (𝓝[>] 0) f (λ t, t ^ (-b))) (hs_bot : b < s.re) : + mellin_convergent f s := begin - rw mellin_convergent_iff_norm (subset_refl _) measurable_set_Ioi - hfc.ae_strongly_measurable, + rw [mellin_convergent, mellin_convergent_iff_norm (subset_refl _) measurable_set_Ioi + hfc.ae_strongly_measurable], exact mellin_convergent_of_is_O_scalar hfc.norm hf_top.norm_left hs_top hf_bot.norm_left hs_bot, end @@ -172,10 +315,10 @@ end /-- If `f` is `O(x ^ (-a))` as `x → 0`, then `log • f` is `O(x ^ (-b))` for every `a < b`. -/ lemma is_O_rpow_zero_log_smul [normed_space ℝ E] {a b : ℝ} {f : ℝ → E} - (hab : a < b) (hf : is_O (𝓝[Ioi 0] 0) f (λ t, t ^ (-a))) : - is_O (𝓝[Ioi 0] 0) (λ t : ℝ, log t • f t) (λ t, t ^ (-b)) := + (hab : a < b) (hf : is_O (𝓝[>] 0) f (λ t, t ^ (-a))) : + is_O (𝓝[>] 0) (λ t : ℝ, log t • f t) (λ t, t ^ (-b)) := begin - have : is_o (𝓝[Ioi 0] 0) log (λ t : ℝ, t ^ (a - b)), + have : is_o (𝓝[>] 0) log (λ t : ℝ, t ^ (a - b)), { refine ((is_o_log_rpow_at_top (sub_pos.mpr hab)).neg_left.comp_tendsto tendsto_inv_zero_at_top).congr' (eventually_nhds_within_iff.mpr $ eventually_of_forall (λ t ht, _)) @@ -198,8 +341,9 @@ theorem mellin_has_deriv_of_is_O_rpow [complete_space E] [normed_space ℂ E] {a b : ℝ} {f : ℝ → E} {s : ℂ} (hfc : locally_integrable_on f $ Ioi 0) (hf_top : is_O at_top f (λ t, t ^ (-a))) (hs_top : s.re < a) - (hf_bot : is_O (𝓝[Ioi 0] 0) f (λ t, t ^ (-b))) (hs_bot : b < s.re) : - has_deriv_at (mellin f) (mellin (λ t, (log t : ℂ) • f t) s) s := + (hf_bot : is_O (𝓝[>] 0) f (λ t, t ^ (-b))) (hs_bot : b < s.re) : + mellin_convergent (λ t, log t • f t) s ∧ + has_deriv_at (mellin f) (mellin (λ t, log t • f t) s) s := begin let F : ℂ → ℝ → E := λ z t, (t : ℂ) ^ (z - 1) • f t, let F' : ℂ → ℝ → E := λ z t, ((t : ℂ) ^ (z - 1) * log t) • f t, @@ -274,8 +418,8 @@ begin rw of_real_log (le_of_lt ht), ring }, exact u1.smul_const (f t) }, - simpa only [F', mellin, mul_smul] using - (has_deriv_at_integral_of_dominated_loc_of_deriv_le hv0 h1 h2 h3 h4 h5 h6).2, + have main := has_deriv_at_integral_of_dominated_loc_of_deriv_le hv0 h1 h2 h3 h4 h5 h6, + exact ⟨by simpa only [F', mul_smul] using main.1, by simpa only [F', mul_smul] using main.2⟩ end /-- Suppose `f` is locally integrable on `(0, ∞)`, is `O(x ^ (-a))` as `x → ∞`, and is @@ -285,9 +429,9 @@ lemma mellin_differentiable_at_of_is_O_rpow [complete_space E] [normed_space ℂ {a b : ℝ} {f : ℝ → E} {s : ℂ} (hfc : locally_integrable_on f $ Ioi 0) (hf_top : is_O at_top f (λ t, t ^ (-a))) (hs_top : s.re < a) - (hf_bot : is_O (𝓝[Ioi 0] 0) f (λ t, t ^ (-b))) (hs_bot : b < s.re) : + (hf_bot : is_O (𝓝[>] 0) f (λ t, t ^ (-b))) (hs_bot : b < s.re) : differentiable_at ℂ (mellin f) s := -(mellin_has_deriv_of_is_O_rpow hfc hf_top hs_top hf_bot hs_bot).differentiable_at +(mellin_has_deriv_of_is_O_rpow hfc hf_top hs_top hf_bot hs_bot).2.differentiable_at end mellin_diff @@ -299,8 +443,8 @@ lemma mellin_convergent_of_is_O_rpow_exp [normed_space ℂ E] {a b : ℝ} (ha : 0 < a) {f : ℝ → E} {s : ℂ} (hfc : locally_integrable_on f $ Ioi 0) (hf_top : is_O at_top f (λ t, exp (-a * t))) - (hf_bot : is_O (𝓝[Ioi 0] 0) f (λ t, t ^ (-b))) (hs_bot : b < s.re) : - integrable_on (λ t : ℝ, (t : ℂ) ^ (s - 1) • f t) (Ioi 0) := + (hf_bot : is_O (𝓝[>] 0) f (λ t, t ^ (-b))) (hs_bot : b < s.re) : + mellin_convergent f s := mellin_convergent_of_is_O_rpow hfc (hf_top.trans (is_o_exp_neg_mul_rpow_at_top ha _).is_O) (lt_add_one _) hf_bot hs_bot @@ -310,9 +454,43 @@ lemma mellin_differentiable_at_of_is_O_rpow_exp [complete_space E] [normed_space {a b : ℝ} (ha : 0 < a) {f : ℝ → E} {s : ℂ} (hfc : locally_integrable_on f $ Ioi 0) (hf_top : is_O at_top f (λ t, exp (-a * t))) - (hf_bot : is_O (𝓝[Ioi 0] 0) f (λ t, t ^ (-b))) (hs_bot : b < s.re) : + (hf_bot : is_O (𝓝[>] 0) f (λ t, t ^ (-b))) (hs_bot : b < s.re) : differentiable_at ℂ (mellin f) s := mellin_differentiable_at_of_is_O_rpow hfc (hf_top.trans (is_o_exp_neg_mul_rpow_at_top ha _).is_O) (lt_add_one _) hf_bot hs_bot end exp_decay + +section mellin_Ioc +/-! +## Mellin transforms of functions on `Ioc 0 1` +-/ + +/-- The Mellin transform of the indicator function of `Ioc 0 1`. -/ +lemma has_mellin_one_Ioc {s : ℂ} (hs : 0 < re s) : + has_mellin (indicator (Ioc 0 1) (λ t, 1 : ℝ → ℂ)) s (1 / s) := +begin + have aux1 : -1 < (s - 1).re, by simpa only [sub_re, one_re, sub_eq_add_neg] + using lt_add_of_pos_left _ hs, + have aux2 : s ≠ 0, by { contrapose! hs, rw [hs, zero_re] }, + have aux3 : measurable_set (Ioc (0 : ℝ) 1), from measurable_set_Ioc, + simp_rw [has_mellin, mellin, mellin_convergent, ←indicator_smul, integrable_on, + integrable_indicator_iff aux3, smul_eq_mul, integral_indicator aux3, + mul_one, integrable_on, measure.restrict_restrict_of_subset Ioc_subset_Ioi_self], + rw [←integrable_on, ←interval_integrable_iff_integrable_Ioc_of_le zero_le_one], + refine ⟨interval_integral.interval_integrable_cpow' aux1, _⟩, + rw [←interval_integral.integral_of_le zero_le_one, integral_cpow (or.inl aux1), sub_add_cancel, + of_real_zero, of_real_one, one_cpow, zero_cpow aux2, sub_zero] +end + +/-- The Mellin transform of a power function restricted to `Ioc 0 1`. -/ +lemma has_mellin_cpow_Ioc (a : ℂ) {s : ℂ} (hs : 0 < re s + re a) : + has_mellin (indicator (Ioc 0 1) (λ t, ↑t ^ a : ℝ → ℂ)) s (1 / (s + a)) := +begin + have := has_mellin_one_Ioc (by rwa add_re : 0 < (s + a).re), + simp_rw [has_mellin, ←mellin_convergent.cpow_smul, ←mellin_cpow_smul, ←indicator_smul, + smul_eq_mul, mul_one] at this, + exact this +end + +end mellin_Ioc diff --git a/src/analysis/special_functions/gamma/basic.lean b/src/analysis/special_functions/gamma/basic.lean index e496adcd4030d..693d12acb5cf7 100644 --- a/src/analysis/special_functions/gamma/basic.lean +++ b/src/analysis/special_functions/gamma/basic.lean @@ -396,7 +396,7 @@ theorem has_deriv_at_Gamma_integral {s : ℂ} (hs : 0 < s.re) : has_deriv_at Gamma_integral (∫ (t : ℝ) in Ioi 0, t ^ (s - 1) * (real.log t * real.exp (-t))) s := begin rw Gamma_integral_eq_mellin, - convert mellin_has_deriv_of_is_O_rpow _ _ (lt_add_one _) _ hs, + convert (mellin_has_deriv_of_is_O_rpow _ _ (lt_add_one _) _ hs).2, { refine (continuous.continuous_on _).locally_integrable_on measurable_set_Ioi, exact continuous_of_real.comp (real.continuous_exp.comp continuous_neg), }, { rw [←is_O_norm_left], From 4f840b8d28320b20c87db17b3a6eef3d325fca87 Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Sun, 28 May 2023 12:56:58 +0000 Subject: [PATCH 1101/1291] feat(ring_theory/polynomial/quotient): generalise quotient_span_X_sub_C_alg_equiv (#19086) --- src/ring_theory/polynomial/quotient.lean | 12 ++++++++++++ 1 file changed, 12 insertions(+) diff --git a/src/ring_theory/polynomial/quotient.lean b/src/ring_theory/polynomial/quotient.lean index 08bb9a5c7ff89..c476511b7c122 100644 --- a/src/ring_theory/polynomial/quotient.lean +++ b/src/ring_theory/polynomial/quotient.lean @@ -37,10 +37,22 @@ rfl (quotient_span_X_sub_C_alg_equiv x).symm y = algebra_map R _ y := rfl +/-- For a commutative ring $R$, evaluating a polynomial at an element $y \in R$ induces an +isomorphism of $R$-algebras $R[X] / \langle x, X - y \rangle \cong R / \langle x \rangle$. -/ +noncomputable def quotient_span_C_X_sub_C_alg_equiv (x y : R) : + (R[X] ⧸ (ideal.span {C x, X - C y} : ideal R[X])) ≃ₐ[R] R ⧸ (ideal.span {x} : ideal R) := +(ideal.quotient_equiv_alg_of_eq R $ by rw [ideal.span_insert, sup_comm]).trans $ + (double_quot.quot_quot_equiv_quot_supₐ R _ _).symm.trans $ + (ideal.quotient_equiv_alg _ _ (quotient_span_X_sub_C_alg_equiv y) rfl).trans $ + ideal.quotient_equiv_alg_of_eq R $ + by { simp only [ideal.map_span, set.image_singleton], congr' 2, exact eval_C } + end polynomial namespace ideal + noncomputable theory + open polynomial variables {R : Type*} [comm_ring R] From d87199d51218d36a0a42c66c82d147b5a7ff87b3 Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Sun, 28 May 2023 14:55:08 +0000 Subject: [PATCH 1102/1291] feat(linear_algebra/free_module/ideal_quotient): add ideal.finrank_quotient_eq_sum (#19084) Co-authored-by: Junyan Xu --- .../free_module/ideal_quotient.lean | 31 +++++++++++++++++-- src/linear_algebra/free_module/pid.lean | 4 +++ 2 files changed, 33 insertions(+), 2 deletions(-) diff --git a/src/linear_algebra/free_module/ideal_quotient.lean b/src/linear_algebra/free_module/ideal_quotient.lean index 7d1ec9fc5f5ef..f0dd7f2c04ceb 100644 --- a/src/linear_algebra/free_module/ideal_quotient.lean +++ b/src/linear_algebra/free_module/ideal_quotient.lean @@ -5,8 +5,9 @@ Authors: Anne Baanen -/ import data.zmod.quotient -import linear_algebra.free_module.finite.basic +import linear_algebra.free_module.finite.rank import linear_algebra.free_module.pid +import linear_algebra.free_module.strong_rank_condition import linear_algebra.quotient_pi /-! # Ideals in free modules over PIDs @@ -21,7 +22,7 @@ import linear_algebra.quotient_pi -/ -open_locale big_operators +open_locale big_operators direct_sum variables {ι R S : Type*} [comm_ring R] [comm_ring S] [algebra R S] variables [is_domain R] [is_principal_ideal_ring R] [is_domain S] [fintype ι] @@ -101,3 +102,29 @@ let b := module.free.choose_basis ℤ S, in by haveI : ∀ i, ne_zero (a i).nat_abs := (λ i, ⟨int.nat_abs_ne_zero_of_ne_zero (ideal.smith_coeffs_ne_zero b I hI i)⟩); classical; exact fintype.of_equiv (Π i, zmod (a i).nat_abs) e.symm + +variables (F : Type*) [comm_ring F] [algebra F R] [algebra F S] [is_scalar_tower F R S] + (b : basis ι R S) {I : ideal S} (hI : I ≠ ⊥) + +/-- Decompose `S⧸I` as a direct sum of cyclic `R`-modules + (quotients by the ideals generated by Smith coefficients of `I`). -/ +noncomputable def ideal.quotient_equiv_direct_sum : + (S ⧸ I) ≃ₗ[F] ⨁ i, R ⧸ ideal.span ({I.smith_coeffs b hI i} : set R) := +begin + apply ((I.quotient_equiv_pi_span b _).restrict_scalars F).trans + (direct_sum.linear_equiv_fun_on_fintype _ _ _).symm, + exact linear_map.is_scalar_tower.compatible_smul + -- why doesn't it automatically apply? + -- even after `change linear_map.compatible_smul _ (Π i, R ⧸ span _) F R` +end + +lemma ideal.finrank_quotient_eq_sum [nontrivial F] + [∀ i, module.free F (R ⧸ ideal.span ({I.smith_coeffs b hI i} : set R))] + [∀ i, module.finite F (R ⧸ ideal.span ({I.smith_coeffs b hI i} : set R))] : + finite_dimensional.finrank F (S ⧸ I) + = ∑ i, finite_dimensional.finrank F (R ⧸ ideal.span ({I.smith_coeffs b hI i} : set R)) := +begin + rw [linear_equiv.finrank_eq $ ideal.quotient_equiv_direct_sum F b hI, + finite_dimensional.finrank_direct_sum] + -- slow, and dot notation doesn't work +end diff --git a/src/linear_algebra/free_module/pid.lean b/src/linear_algebra/free_module/pid.lean index af76def8e465f..a54dbec23e0bb 100644 --- a/src/linear_algebra/free_module/pid.lean +++ b/src/linear_algebra/free_module/pid.lean @@ -608,6 +608,10 @@ begin simp [hi] end +instance (F : Type u) [comm_ring F] [algebra F R] (b : basis ι R S) {I : ideal S} (hI : I ≠ ⊥) (i) : + module F (R ⧸ ideal.span ({I.smith_coeffs b hI i} : set R)) := +by apply_instance -- quotient.module' _ + end ideal end smith_normal From f7ebde7ee0d1505dfccac8644ae12371aa3c1c9f Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Sun, 28 May 2023 14:55:09 +0000 Subject: [PATCH 1103/1291] refactor(topology/vector_bundle/hom): fibres of hom-bundle carry strong topology (#19107) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Currently, the "hom-bundle" between two vector bundles `E₁` and `E₂` has fibre over `x` which is a type synonym of `E₁ x →SL[σ] E₂ x`, but which carries a topology produced by the hom-bundle construction (using the identification by trivializations withe the model fibre `F₁ →SL[σ] F₂`). This was needed when this bundle was made (#14541) because at that time, `F₁ →SL[σ] F₂` (continuous linear maps between normed spaces) carried a topology in mathlib but `E₁ x →SL[σ] E₂ x` (continuous linear maps between topological vector spaces) did not. As of #16053, continuous linear maps between topological vector spaces do carry a topology, the strong topology. So we can kill the old topology on the type synonym and just use the default one, which should avoid annoying issues later. A few minor changes are needed to make this go through: - we revert #14377: the question is whether the "vector prebundle" construction, whose canonical use is for the hom-bundle, should or should not require a topology on the fibres. Now that in applications it could happen either way (fibres do or don't come with a topology), it will be more convenient to assume that they do carry a topology, and put the "artificial" topology on the fibres if they happen to not. - some assumptions need to change from `[add_comm_monoid]` to `[add_comm_group]`, this is mathematically harmless since they are also modules over a field. - generalize the construction `continuous_linear_equiv.arrow_congrSL` from normed spaces to topological vector spaces Co-authored-by: Moritz Doll Co-authored-by: Floris van Doorn --- src/analysis/convolution.lean | 4 +- src/analysis/normed_space/operator_norm.lean | 40 --------- .../manifold/vector_bundle/basic.lean | 4 +- src/geometry/manifold/vector_bundle/hom.lean | 21 ++--- .../algebra/module/strong_topology.lean | 83 +++++++++++++++++ src/topology/constructions.lean | 8 ++ src/topology/fiber_bundle/basic.lean | 38 +++++--- src/topology/vector_bundle/basic.lean | 19 ++-- src/topology/vector_bundle/hom.lean | 90 ++++++++++++------- 9 files changed, 192 insertions(+), 115 deletions(-) diff --git a/src/analysis/convolution.lean b/src/analysis/convolution.lean index 8d79eeab2c1a9..edcbfe3a605bb 100644 --- a/src/analysis/convolution.lean +++ b/src/analysis/convolution.lean @@ -1497,7 +1497,9 @@ begin simp only [ef, eg, comp_app, continuous_linear_equiv.apply_symm_apply, coe_comp', continuous_linear_equiv.prod_apply, continuous_linear_equiv.map_sub, continuous_linear_equiv.arrow_congr, continuous_linear_equiv.arrow_congrSL_symm_apply, - continuous_linear_equiv.coe_coe, comp_app, continuous_linear_equiv.apply_symm_apply] }, + continuous_linear_equiv.coe_coe, comp_app, continuous_linear_equiv.apply_symm_apply, + linear_equiv.inv_fun_eq_symm, continuous_linear_equiv.arrow_congrₛₗ_symm_apply, + eq_self_iff_true] }, simp_rw [this] at A, exact A, end diff --git a/src/analysis/normed_space/operator_norm.lean b/src/analysis/normed_space/operator_norm.lean index f497c31dbaba2..6b39d840c5f70 100644 --- a/src/analysis/normed_space/operator_norm.lean +++ b/src/analysis/normed_space/operator_norm.lean @@ -1715,46 +1715,6 @@ begin (λ y, homothety_inverse _ hx _ (to_span_nonzero_singleton_homothety 𝕜 x h) _) end -variables {𝕜} {𝕜₄ : Type*} [nontrivially_normed_field 𝕜₄] -variables {H : Type*} [normed_add_comm_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. -/ -@[simps apply symm_apply] -def arrow_congrSL (e₁₂ : E ≃SL[σ₁₂] F) (e₄₃ : H ≃SL[σ₄₃] G) : - (E →SL[σ₁₄] H) ≃SL[σ₄₃] (F →SL[σ₂₃] G) := -{ -- given explicitly to help `simps` - to_fun := λ L, (e₄₃ : H →SL[σ₄₃] G).comp (L.comp (e₁₂.symm : F →SL[σ₂₁] E)), - -- given explicitly to help `simps` - inv_fun := λ L, (e₄₃.symm : G →SL[σ₃₄] H).comp (L.comp (e₁₂ : E →SL[σ₁₂] F)), - map_add' := λ f g, by rw [add_comp, comp_add], - map_smul' := λ t f, by rw [smul_comp, comp_smulₛₗ], - continuous_to_fun := (continuous_id.clm_comp_const _).const_clm_comp _, - continuous_inv_fun := (continuous_id.clm_comp_const _).const_clm_comp _, - .. 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_add_comm_group F] [normed_add_comm_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/geometry/manifold/vector_bundle/basic.lean b/src/geometry/manifold/vector_bundle/basic.lean index 0b38987317863..01fc29ef7ebf3 100644 --- a/src/geometry/manifold/vector_bundle/basic.lean +++ b/src/geometry/manifold/vector_bundle/basic.lean @@ -394,7 +394,7 @@ end with_topology namespace vector_prebundle -variables {F E} +variables [∀ x, topological_space (E x)] {F E} /-- Mixin for a `vector_prebundle` stating smoothness of coordinate changes. -/ class is_smooth (a : vector_prebundle 𝕜 F E) : Prop := @@ -438,7 +438,7 @@ variables (IB) /-- Make a `smooth_vector_bundle` from a `smooth_vector_prebundle`. -/ lemma smooth_vector_bundle : @smooth_vector_bundle _ _ F E _ _ _ _ _ _ IB _ _ _ _ _ _ _ - a.total_space_topology a.fiber_topology a.to_fiber_bundle a.to_vector_bundle := + a.total_space_topology _ a.to_fiber_bundle a.to_vector_bundle := { smooth_on_coord_change := begin rintros _ _ ⟨e, he, rfl⟩ ⟨e', he', rfl⟩, refine (a.smooth_on_smooth_coord_change he he').congr _, diff --git a/src/geometry/manifold/vector_bundle/hom.lean b/src/geometry/manifold/vector_bundle/hom.lean index d81bcf54cbe50..e0930af6bb21f 100644 --- a/src/geometry/manifold/vector_bundle/hom.lean +++ b/src/geometry/manifold/vector_bundle/hom.lean @@ -23,15 +23,16 @@ open_locale manifold bundle variables {𝕜 B F F₁ F₂ M M₁ M₂ : Type*} {E : B → Type*} {E₁ : B → Type*} {E₂ : B → Type*} [nontrivially_normed_field 𝕜] - [∀ x, add_comm_monoid (E x)] [∀ x, module 𝕜 (E x)] + [∀ x, add_comm_group (E x)] [∀ x, module 𝕜 (E x)] [normed_add_comm_group F] [normed_space 𝕜 F] [topological_space (total_space E)] [∀ x, topological_space (E x)] - [∀ x, add_comm_monoid (E₁ x)] [∀ x, module 𝕜 (E₁ x)] + [∀ x, add_comm_group (E₁ x)] [∀ x, module 𝕜 (E₁ x)] [normed_add_comm_group F₁] [normed_space 𝕜 F₁] [topological_space (total_space E₁)] [∀ x, topological_space (E₁ x)] - [∀ x, add_comm_monoid (E₂ x)] [∀ x, module 𝕜 (E₂ x)] + [∀ x, add_comm_group (E₂ x)] [∀ x, module 𝕜 (E₂ x)] [normed_add_comm_group F₂] [normed_space 𝕜 F₂] [topological_space (total_space E₂)] [∀ x, topological_space (E₂ x)] + [_i₁ : ∀ x, topological_add_group (E₂ x)] [_i₂ : ∀ x, has_continuous_smul 𝕜 (E₂ x)] {EB : Type*} [normed_add_comm_group EB] [normed_space 𝕜 EB] {HB : Type*} [topological_space HB] (IB : model_with_corners 𝕜 EB HB) @@ -67,10 +68,11 @@ begin { intros b hb, ext L v, simp only [continuous_linear_map_coord_change, continuous_linear_equiv.coe_coe, continuous_linear_equiv.arrow_congrSL_apply, comp_apply, function.comp, compL_apply, - flip_apply, continuous_linear_equiv.symm_symm] }, + flip_apply, continuous_linear_equiv.symm_symm, linear_equiv.to_fun_eq_coe, + continuous_linear_equiv.arrow_congrₛₗ_apply, continuous_linear_map.coe_comp'] }, end -variables [∀ x, has_continuous_add (E₂ x)] [∀ x, has_continuous_smul 𝕜 (E₂ x)] +include _i₁ _i₂ lemma hom_chart (y₀ y : LE₁E₂) : chart_at (model_prod HB (F₁ →L[𝕜] F₂)) y₀ y = @@ -107,15 +109,6 @@ instance bundle.continuous_linear_map.vector_prebundle.is_smooth : continuous_linear_map_coord_change_apply (ring_hom.id 𝕜) e₁ e₁' e₂ e₂'⟩ end } -/-- Todo: remove this definition. It is probably needed because of the type-class pi bug -https://leanprover.zulipchat.com/#narrow/stream/116395-maths/topic/vector.20bundles.20--.20typeclass.20inference.20issue --/ -@[reducible] -def smooth_vector_bundle.continuous_linear_map.aux (x) : - topological_space (bundle.continuous_linear_map (ring_hom.id 𝕜) F₁ E₁ F₂ E₂ x) := -by apply_instance -local attribute [instance, priority 1] smooth_vector_bundle.continuous_linear_map.aux - instance smooth_vector_bundle.continuous_linear_map : smooth_vector_bundle (F₁ →L[𝕜] F₂) (bundle.continuous_linear_map (ring_hom.id 𝕜) F₁ E₁ F₂ E₂) IB := diff --git a/src/topology/algebra/module/strong_topology.lean b/src/topology/algebra/module/strong_topology.lean index a1ad2e299b21f..706f6818c70fe 100644 --- a/src/topology/algebra/module/strong_topology.lean +++ b/src/topology/algebra/module/strong_topology.lean @@ -227,3 +227,86 @@ continuous_linear_map.has_basis_nhds_zero_of_basis (𝓝 0).basis_sets end bounded_sets end continuous_linear_map + +open continuous_linear_map + +namespace continuous_linear_equiv + +section semilinear + +variables {𝕜 : Type*} {𝕜₂ : Type*} {𝕜₃ : Type*} {𝕜₄ : Type*} + {E : Type*} {F : Type*} {G : Type*} {H : Type*} + [add_comm_group E] [add_comm_group F] [add_comm_group G] [add_comm_group H] + [nontrivially_normed_field 𝕜] [nontrivially_normed_field 𝕜₂] [nontrivially_normed_field 𝕜₃] + [nontrivially_normed_field 𝕜₄] + [module 𝕜 E] [module 𝕜₂ F] [module 𝕜₃ G] [module 𝕜₄ H] + [topological_space E] [topological_space F] [topological_space G] [topological_space H] + [topological_add_group G] [topological_add_group H] + [has_continuous_const_smul 𝕜₃ G] [has_continuous_const_smul 𝕜₄ H] + {σ₁₂ : 𝕜 →+* 𝕜₂} {σ₂₁ : 𝕜₂ →+* 𝕜} {σ₂₃ : 𝕜₂ →+* 𝕜₃} {σ₁₃ : 𝕜 →+* 𝕜₃} {σ₃₄ : 𝕜₃ →+* 𝕜₄} + {σ₄₃ : 𝕜₄ →+* 𝕜₃} {σ₂₄ : 𝕜₂ →+* 𝕜₄} {σ₁₄ : 𝕜 →+* 𝕜₄} + [ring_hom_inv_pair σ₁₂ σ₂₁] [ring_hom_inv_pair σ₂₁ σ₁₂] [ring_hom_inv_pair σ₃₄ σ₄₃] + [ring_hom_inv_pair σ₄₃ σ₃₄] + [ring_hom_comp_triple σ₂₁ σ₁₄ σ₂₄] [ring_hom_comp_triple σ₂₄ σ₄₃ σ₂₃] + [ring_hom_comp_triple σ₁₂ σ₂₃ σ₁₃] [ring_hom_comp_triple σ₁₃ σ₃₄ σ₁₄] + +include σ₁₄ σ₂₄ σ₁₃ σ₃₄ σ₂₁ σ₂₃ + +/-- A pair of continuous (semi)linear equivalences generates a (semi)linear equivalence between the +spaces of continuous (semi)linear maps. -/ +@[simps] def arrow_congrₛₗ (e₁₂ : E ≃SL[σ₁₂] F) (e₄₃ : H ≃SL[σ₄₃] G) : + (E →SL[σ₁₄] H) ≃ₛₗ[σ₄₃] (F →SL[σ₂₃] G) := +{ -- given explicitly to help `simps` + to_fun := λ L, (e₄₃ : H →SL[σ₄₃] G).comp (L.comp (e₁₂.symm : F →SL[σ₂₁] E)), + -- given explicitly to help `simps` + inv_fun := λ L, (e₄₃.symm : G →SL[σ₃₄] H).comp (L.comp (e₁₂ : E →SL[σ₁₂] F)), + map_add' := λ f g, by rw [add_comp, comp_add], + map_smul' := λ t f, by rw [smul_comp, comp_smulₛₗ], + .. e₁₂.arrow_congr_equiv e₄₃, } + +variables [ring_hom_isometric σ₂₁] + +lemma arrow_congrₛₗ_continuous (e₁₂ : E ≃SL[σ₁₂] F) (e₄₃ : H ≃SL[σ₄₃] G) : + continuous (id (e₁₂.arrow_congrₛₗ e₄₃ : (E →SL[σ₁₄] H) ≃ₛₗ[σ₄₃] (F →SL[σ₂₃] G))) := +begin + apply continuous_of_continuous_at_zero, + show filter.tendsto _ _ _, + simp_rw [(e₁₂.arrow_congrₛₗ e₄₃).map_zero], + rw continuous_linear_map.has_basis_nhds_zero.tendsto_iff + continuous_linear_map.has_basis_nhds_zero, + rintros ⟨sF, sG⟩ ⟨h1 : bornology.is_vonN_bounded 𝕜₂ sF, h2 : sG ∈ nhds (0:G)⟩, + dsimp, + refine ⟨(e₁₂.symm '' sF, e₄₃ ⁻¹' sG), ⟨h1.image (e₁₂.symm : F →SL[σ₂₁] E), _⟩, + λ _ h _ hx, h _ (set.mem_image_of_mem _ hx)⟩, + apply e₄₃.continuous.continuous_at, + simpa using h2, +end + +variables [ring_hom_isometric σ₁₂] + +/-- A pair of continuous (semi)linear equivalences generates an continuous (semi)linear equivalence +between the spaces of continuous (semi)linear maps. -/ +@[simps] def arrow_congrSL (e₁₂ : E ≃SL[σ₁₂] F) (e₄₃ : H ≃SL[σ₄₃] G) : + (E →SL[σ₁₄] H) ≃SL[σ₄₃] (F →SL[σ₂₃] G) := +{ continuous_to_fun := e₁₂.arrow_congrₛₗ_continuous e₄₃, + continuous_inv_fun := e₁₂.symm.arrow_congrₛₗ_continuous e₄₃.symm, + .. e₁₂.arrow_congrₛₗ e₄₃, } + +end semilinear + +section linear +variables {𝕜 : Type*} {E : Type*} {F : Type*} {G : Type*} {H : Type*} + [add_comm_group E] [add_comm_group F] [add_comm_group G] [add_comm_group H] + [nontrivially_normed_field 𝕜] [module 𝕜 E] [module 𝕜 F] [module 𝕜 G] [module 𝕜 H] + [topological_space E] [topological_space F] [topological_space G] [topological_space H] + [topological_add_group G] [topological_add_group H] + [has_continuous_const_smul 𝕜 G] [has_continuous_const_smul 𝕜 H] + +/-- A pair of continuous linear equivalences generates an continuous linear equivalence between +the spaces of continuous linear maps. -/ +def arrow_congr (e₁ : E ≃L[𝕜] F) (e₂ : H ≃L[𝕜] G) : (E →L[𝕜] H) ≃L[𝕜] (F →L[𝕜] G) := +e₁.arrow_congrSL e₂ + +end linear + +end continuous_linear_equiv diff --git a/src/topology/constructions.lean b/src/topology/constructions.lean index f993d015b760f..381b14ffe7d47 100644 --- a/src/topology/constructions.lean +++ b/src/topology/constructions.lean @@ -680,6 +680,14 @@ lemma inducing.prod_mk {f : α → β} {g : γ → δ} (hf : inducing f) (hg : i ⟨by rw [prod.topological_space, prod.topological_space, hf.induced, hg.induced, induced_compose, induced_compose, induced_inf, induced_compose, induced_compose]⟩ +@[simp] lemma inducing_const_prod {a : α} {f : β → γ} : inducing (λ x, (a, f x)) ↔ inducing f := +by simp_rw [inducing_iff, prod.topological_space, induced_inf, induced_compose, function.comp, + induced_const, top_inf_eq] + +@[simp] lemma inducing_prod_const {b : β} {f : α → γ} : inducing (λ x, (f x, b)) ↔ inducing f := +by simp_rw [inducing_iff, prod.topological_space, induced_inf, induced_compose, function.comp, + induced_const, inf_top_eq] + lemma embedding.prod_mk {f : α → β} {g : γ → δ} (hf : embedding f) (hg : embedding g) : embedding (λx:α×γ, (f x.1, g x.2)) := { inj := assume ⟨x₁, x₂⟩ ⟨y₁, y₂⟩, by simp; exact assume h₁ h₂, ⟨hf.inj h₁, hg.inj h₂⟩, diff --git a/src/topology/fiber_bundle/basic.lean b/src/topology/fiber_bundle/basic.lean index f9271f8899861..f50b7b4f4e649 100644 --- a/src/topology/fiber_bundle/basic.lean +++ b/src/topology/fiber_bundle/basic.lean @@ -705,6 +705,7 @@ end fiber_bundle_core /-! ### Prebundle construction for constructing fiber bundles -/ variables (F) (E : B → Type*) [topological_space B] [topological_space F] + [Π x, topological_space (E x)] /-- This structure permits to define a fiber bundle when trivializations are given as local equivalences but there is not yet a topology on the total space. The total space is hence given a @@ -718,6 +719,7 @@ structure fiber_prebundle := (pretrivialization_mem_atlas : ∀ x : B, pretrivialization_at x ∈ pretrivialization_atlas) (continuous_triv_change : ∀ e e' ∈ pretrivialization_atlas, continuous_on (e ∘ e'.to_local_equiv.symm) (e'.target ∩ (e'.to_local_equiv.symm ⁻¹' e.source))) +(total_space_mk_inducing : ∀ (b : B), inducing ((pretrivialization_at b) ∘ (total_space_mk b))) namespace fiber_prebundle @@ -799,19 +801,27 @@ begin exact a.mem_base_pretrivialization_at b, end -/-- Topology on the fibers `E b` induced by the map `E b → E.total_space`. -/ -def fiber_topology (b : B) : topological_space (E b) := -topological_space.induced (total_space_mk b) a.total_space_topology - -@[continuity] lemma inducing_total_space_mk (b : B) : - @inducing _ _ (a.fiber_topology b) a.total_space_topology (total_space_mk b) := -by { letI := a.total_space_topology, letI := a.fiber_topology b, exact ⟨rfl⟩ } - @[continuity] lemma continuous_total_space_mk (b : B) : - @continuous _ _ (a.fiber_topology b) a.total_space_topology (total_space_mk b) := + @continuous _ _ _ a.total_space_topology (total_space_mk b) := begin - letI := a.total_space_topology, letI := a.fiber_topology b, - exact (a.inducing_total_space_mk b).continuous + letI := a.total_space_topology, + let e := a.trivialization_of_mem_pretrivialization_atlas (a.pretrivialization_mem_atlas b), + rw e.to_local_homeomorph.continuous_iff_continuous_comp_left + (a.total_space_mk_preimage_source b), + exact continuous_iff_le_induced.mpr (le_antisymm_iff.mp (a.total_space_mk_inducing b).induced).1, +end + +lemma inducing_total_space_mk_of_inducing_comp (b : B) + (h : inducing ((a.pretrivialization_at b) ∘ (total_space_mk b))) : + @inducing _ _ _ a.total_space_topology (total_space_mk b) := +begin + letI := a.total_space_topology, + rw ←restrict_comp_cod_restrict (a.mem_trivialization_at_source b) at h, + apply inducing.of_cod_restrict (a.mem_trivialization_at_source b), + refine inducing_of_inducing_compose _ (continuous_on_iff_continuous_restrict.mp + (a.trivialization_of_mem_pretrivialization_atlas + (a.pretrivialization_mem_atlas b)).continuous_to_fun) h, + exact (a.continuous_total_space_mk b).cod_restrict (a.mem_trivialization_at_source b), end /-- Make a `fiber_bundle` from a `fiber_prebundle`. Concretely this means @@ -821,8 +831,9 @@ establishes that for the topology constructed on the sigma-type using `fiber_prebundle.total_space_topology`, these "pretrivializations" are actually "trivializations" (i.e., homeomorphisms with respect to the constructed topology). -/ def to_fiber_bundle : - @fiber_bundle B F _ _ E a.total_space_topology a.fiber_topology := -{ total_space_mk_inducing := a.inducing_total_space_mk, + @fiber_bundle B F _ _ E a.total_space_topology _ := +{ total_space_mk_inducing := λ b, a.inducing_total_space_mk_of_inducing_comp b + (a.total_space_mk_inducing b), trivialization_atlas := {e | ∃ e₀ (he₀ : e₀ ∈ a.pretrivialization_atlas), e = a.trivialization_of_mem_pretrivialization_atlas he₀}, trivialization_at := λ x, a.trivialization_of_mem_pretrivialization_atlas @@ -833,7 +844,6 @@ def to_fiber_bundle : lemma continuous_proj : @continuous _ _ a.total_space_topology _ (π E) := begin letI := a.total_space_topology, - letI := a.fiber_topology, letI := a.to_fiber_bundle, exact continuous_proj F E, end diff --git a/src/topology/vector_bundle/basic.lean b/src/topology/vector_bundle/basic.lean index 483fa7312e6ba..6d1d563cfcd6c 100644 --- a/src/topology/vector_bundle/basic.lean +++ b/src/topology/vector_bundle/basic.lean @@ -733,7 +733,7 @@ end section variables [nontrivially_normed_field R] [∀ x, add_comm_monoid (E x)] [∀ x, module R (E x)] - [normed_add_comm_group F] [normed_space R F] [topological_space B] + [normed_add_comm_group F] [normed_space R F] [topological_space B] [∀ x, topological_space (E x)] open topological_space @@ -760,6 +760,7 @@ structure vector_prebundle := continuous_on f (e.base_set ∩ e'.base_set) ∧ ∀ (b : B) (hb : b ∈ e.base_set ∩ e'.base_set) (v : F), f b v = (e' (total_space_mk b (e.symm b v))).2) +(total_space_mk_inducing : ∀ (b : B), inducing ((pretrivialization_at b) ∘ (total_space_mk b))) namespace vector_prebundle @@ -845,21 +846,13 @@ a.to_fiber_prebundle.mem_trivialization_at_source b x (total_space_mk b) ⁻¹' (a.pretrivialization_at b).source = univ := a.to_fiber_prebundle.total_space_mk_preimage_source b -/-- Topology on the fibers `E b` induced by the map `E b → E.total_space`. -/ -def fiber_topology (b : B) : topological_space (E b) := -a.to_fiber_prebundle.fiber_topology b - -@[continuity] lemma inducing_total_space_mk (b : B) : - @inducing _ _ (a.fiber_topology b) a.total_space_topology (total_space_mk b) := -a.to_fiber_prebundle.inducing_total_space_mk b - @[continuity] lemma continuous_total_space_mk (b : B) : - @continuous _ _ (a.fiber_topology b) a.total_space_topology (total_space_mk b) := + @continuous _ _ _ a.total_space_topology (total_space_mk b) := a.to_fiber_prebundle.continuous_total_space_mk b /-- Make a `fiber_bundle` from a `vector_prebundle`; auxiliary construction for `vector_prebundle.vector_bundle`. -/ -def to_fiber_bundle : @fiber_bundle B F _ _ _ a.total_space_topology a.fiber_topology := +def to_fiber_bundle : @fiber_bundle B F _ _ _ a.total_space_topology _ := a.to_fiber_prebundle.to_fiber_bundle /-- Make a `vector_bundle` from a `vector_prebundle`. Concretely this means @@ -869,7 +862,7 @@ establishes that for the topology constructed on the sigma-type using `vector_prebundle.total_space_topology`, these "pretrivializations" are actually "trivializations" (i.e., homeomorphisms with respect to the constructed topology). -/ lemma to_vector_bundle : - @vector_bundle R _ F E _ _ _ _ _ _ a.total_space_topology a.fiber_topology a.to_fiber_bundle := + @vector_bundle R _ F E _ _ _ _ _ _ a.total_space_topology _ a.to_fiber_bundle := { trivialization_linear' := begin rintros _ ⟨e, he, rfl⟩, apply linear_of_mem_pretrivialization_atlas, @@ -895,7 +888,7 @@ variables [normed_space 𝕜₁ F] [Π x, module 𝕜₁ (E x)] [topological_spa variables {F' : Type*} [normed_add_comm_group F'] [normed_space 𝕜₂ F'] {E' : B' → Type*} [Π x, add_comm_monoid (E' x)] [Π x, module 𝕜₂ (E' x)] [topological_space (total_space E')] -variables [Π x, topological_space (E x)] [fiber_bundle F E] [vector_bundle 𝕜₁ F E] +variables [fiber_bundle F E] [vector_bundle 𝕜₁ F E] variables [Π x, topological_space (E' x)] [fiber_bundle F' E'] [vector_bundle 𝕜₂ F' E'] variables (F E F' E') diff --git a/src/topology/vector_bundle/hom.lean b/src/topology/vector_bundle/hom.lean index 20c00d4d11894..68258be0aa117 100644 --- a/src/topology/vector_bundle/hom.lean +++ b/src/topology/vector_bundle/hom.lean @@ -18,11 +18,11 @@ their respective scalar fields, we define `bundle.continuous_linear_map σ F₁ type synonym for `λ x, E₁ x →SL[σ] E₂ x`. If the `E₁` and `E₂` are vector bundles with model fibers `F₁` and `F₂`, then this will be a vector bundle with fiber `F₁ →SL[σ] F₂`. -The topology is constructed from the trivializations for `E₁` and `E₂` and the norm-topology on the -model fiber `F₁ →SL[𝕜] F₂` using the `vector_prebundle` construction. This is a bit awkward because -it introduces a spurious (?) dependence on the normed space structure of the model fiber, rather -than just its topological vector space structure; this might be fixable now that we have -`continuous_linear_map.strong_topology`. +The topology on the total space is constructed from the trivializations for `E₁` and `E₂` and the +norm-topology on the model fiber `F₁ →SL[𝕜] F₂` using the `vector_prebundle` construction. This is +a bit awkward because it introduces a dependence on the normed space structure of the model fibers, +rather than just their topological vector space structure; it is not clear whether this is +necessary. Similar constructions should be possible (but are yet to be formalized) for tensor products of topological vector bundles, exterior algebras, and so on, where again the topology can be defined @@ -44,9 +44,9 @@ section defs variables {𝕜₁ 𝕜₂ : Type*} [normed_field 𝕜₁] [normed_field 𝕜₂] variables (σ : 𝕜₁ →+* 𝕜₂) variables {B : Type*} -variables (F₁ : Type*) (E₁ : B → Type*) [Π x, add_comm_monoid (E₁ x)] [Π x, module 𝕜₁ (E₁ x)] +variables (F₁ : Type*) (E₁ : B → Type*) [Π x, add_comm_group (E₁ x)] [Π x, module 𝕜₁ (E₁ x)] variables [Π x, topological_space (E₁ x)] -variables (F₂ : Type*) (E₂ : B → Type*) [Π x, add_comm_monoid (E₂ x)] [Π x, module 𝕜₂ (E₂ x)] +variables (F₂ : Type*) (E₂ : B → Type*) [Π x, add_comm_group (E₂ x)] [Π x, module 𝕜₂ (E₂ x)] variables [Π x, topological_space (E₂ x)] include F₁ F₂ @@ -68,7 +68,10 @@ instance bundle.continuous_linear_map.add_monoid_hom_class (x : B) : add_monoid_hom_class (bundle.continuous_linear_map σ F₁ E₁ F₂ E₂ x) (E₁ x) (E₂ x) := by delta_instance bundle.continuous_linear_map -variables [Π x, has_continuous_add (E₂ x)] +variables [Π x, topological_add_group (E₂ x)] + +instance (x : B) : topological_space (bundle.continuous_linear_map σ F₁ E₁ F₂ E₂ x) := +by delta_instance bundle.continuous_linear_map instance (x : B) : add_comm_monoid (bundle.continuous_linear_map σ F₁ E₁ F₂ E₂ x) := by delta_instance bundle.continuous_linear_map @@ -86,18 +89,16 @@ variables {𝕜₁ : Type*} [nontrivially_normed_field 𝕜₁] {𝕜₂ : Type* variables {B : Type*} [topological_space B] variables (F₁ : Type*) [normed_add_comm_group F₁] [normed_space 𝕜₁ F₁] - (E₁ : B → Type*) [Π x, add_comm_monoid (E₁ x)] [Π x, module 𝕜₁ (E₁ x)] + (E₁ : B → Type*) [Π x, add_comm_group (E₁ x)] [Π x, module 𝕜₁ (E₁ x)] [topological_space (total_space E₁)] variables (F₂ : Type*) [normed_add_comm_group F₂][normed_space 𝕜₂ F₂] - (E₂ : B → Type*) [Π x, add_comm_monoid (E₂ x)] [Π x, module 𝕜₂ (E₂ x)] + (E₂ : B → Type*) [Π x, add_comm_group (E₂ x)] [Π x, module 𝕜₂ (E₂ x)] [topological_space (total_space E₂)] variables {F₁ E₁ F₂ E₂} (e₁ e₁' : trivialization F₁ (π E₁)) (e₂ e₂' : trivialization F₂ (π E₂)) namespace pretrivialization -include iσ - /-- Assume `eᵢ` and `eᵢ'` are trivializations of the bundles `Eᵢ` over base `B` with fiber `Fᵢ` (`i ∈ {1,2}`), then `continuous_linear_map_coord_change σ e₁ e₁' e₂ e₂'` is the coordinate change function between the two induced (pre)trivializations @@ -111,7 +112,10 @@ def continuous_linear_map_coord_change variables {σ e₁ e₁' e₂ e₂'} variables [Π x, topological_space (E₁ x)] [fiber_bundle F₁ E₁] -variables [Π x, topological_space (E₂ x)] [fiber_bundle F₂ E₂] +variables [Π x, topological_space (E₂ x)] [ita : Π x, topological_add_group (E₂ x)] + [fiber_bundle F₂ E₂] + +include iσ lemma continuous_on_continuous_linear_map_coord_change [vector_bundle 𝕜₁ F₁ E₁] [vector_bundle 𝕜₂ F₂ E₂] @@ -129,6 +133,7 @@ begin { mfld_set_tac }, { intros b hb, ext L v, simp only [continuous_linear_map_coord_change, continuous_linear_equiv.coe_coe, + continuous_linear_equiv.arrow_congrₛₗ_apply, linear_equiv.to_fun_eq_coe, coe_comp', continuous_linear_equiv.arrow_congrSL_apply, comp_apply, function.comp, compSL_apply, flip_apply, continuous_linear_equiv.symm_symm] }, end @@ -145,8 +150,10 @@ trivialization, after the bundle of continuous semilinear maps is equipped with topological vector bundle structure. -/ def continuous_linear_map : pretrivialization (F₁ →SL[σ] F₂) (π (bundle.continuous_linear_map σ F₁ E₁ F₂ E₂)) := -{ to_fun := λ p, ⟨p.1, (e₂.continuous_linear_map_at 𝕜₂ p.1).comp $ p.2.comp $ e₁.symmL 𝕜₁ p.1⟩, - inv_fun := λ p, ⟨p.1, (e₂.symmL 𝕜₂ p.1).comp $ p.2.comp $ e₁.continuous_linear_map_at 𝕜₁ p.1⟩, +{ to_fun := λ p, ⟨p.1, continuous_linear_map.comp (e₂.continuous_linear_map_at 𝕜₂ p.1) + (p.2.comp (e₁.symmL 𝕜₁ p.1 : F₁ →L[𝕜₁] E₁ p.1) : F₁ →SL[σ] E₂ p.1)⟩, + inv_fun := λ p, ⟨p.1, continuous_linear_map.comp (e₂.symmL 𝕜₂ p.1) + (p.2.comp (e₁.continuous_linear_map_at 𝕜₁ p.1 : E₁ p.1 →L[𝕜₁] F₁) : E₁ p.1 →SL[σ] F₂)⟩, source := (bundle.total_space.proj) ⁻¹' (e₁.base_set ∩ e₂.base_set), target := (e₁.base_set ∩ e₂.base_set) ×ˢ set.univ, map_source' := λ ⟨x, L⟩ h, ⟨h, set.mem_univ _⟩, @@ -170,6 +177,8 @@ def continuous_linear_map : target_eq := rfl, proj_to_fun := λ ⟨x, f⟩ h, rfl } +include ita + instance continuous_linear_map.is_linear [Π x, has_continuous_add (E₂ x)] [Π x, has_continuous_smul 𝕜₂ (E₂ x)] : (pretrivialization.continuous_linear_map σ e₁ e₂).is_linear 𝕜₂ := @@ -187,18 +196,22 @@ instance continuous_linear_map.is_linear refl end, } } +omit ita + lemma continuous_linear_map_apply (p : total_space (bundle.continuous_linear_map σ F₁ E₁ F₂ E₂)) : (continuous_linear_map σ e₁ e₂) p = - ⟨p.1, (e₂.continuous_linear_map_at 𝕜₂ p.1).comp $ p.2.comp $ e₁.symmL 𝕜₁ p.1⟩ := + ⟨p.1, continuous_linear_map.comp (e₂.continuous_linear_map_at 𝕜₂ p.1) + (p.2.comp (e₁.symmL 𝕜₁ p.1 : F₁ →L[𝕜₁] E₁ p.1) : F₁ →SL[σ] E₂ p.1)⟩ := rfl lemma continuous_linear_map_symm_apply (p : B × (F₁ →SL[σ] F₂)) : (continuous_linear_map σ e₁ e₂).to_local_equiv.symm p = - ⟨p.1, (e₂.symmL 𝕜₂ p.1).comp $ p.2.comp $ e₁.continuous_linear_map_at 𝕜₁ p.1⟩ := + ⟨p.1, continuous_linear_map.comp (e₂.symmL 𝕜₂ p.1) + (p.2.comp (e₁.continuous_linear_map_at 𝕜₁ p.1 : E₁ p.1 →L[𝕜₁] F₁) : E₁ p.1 →SL[σ] F₂)⟩ := rfl -variables [Π x, has_continuous_add (E₂ x)] +include ita lemma continuous_linear_map_symm_apply' {b : B} (hb : b ∈ e₁.base_set ∩ e₂.base_set) (L : F₁ →SL[σ] F₂) : @@ -208,7 +221,7 @@ begin rw [symm_apply], refl, exact hb end -lemma continuous_linear_map_coord_change_apply [ring_hom_isometric σ] (b : B) +lemma continuous_linear_map_coord_change_apply (b : B) (hb : b ∈ (e₁.base_set ∩ e₂.base_set) ∩ (e₁'.base_set ∩ e₂'.base_set)) (L : F₁ →SL[σ] F₂) : continuous_linear_map_coord_change σ e₁ e₁' e₂ e₂' b L = (continuous_linear_map σ e₁' e₂' @@ -216,7 +229,8 @@ lemma continuous_linear_map_coord_change_apply [ring_hom_isometric σ] (b : B) begin ext v, simp_rw [continuous_linear_map_coord_change, continuous_linear_equiv.coe_coe, - continuous_linear_equiv.arrow_congrSL_apply, + continuous_linear_equiv.arrow_congrSL_apply, linear_equiv.to_fun_eq_coe, + continuous_linear_equiv.arrow_congrₛₗ_apply, continuous_linear_map_apply, continuous_linear_map_symm_apply' σ e₁ e₂ hb.1, comp_apply, continuous_linear_equiv.coe_coe, continuous_linear_equiv.symm_symm, trivialization.continuous_linear_map_at_apply, trivialization.symmL_apply], @@ -229,10 +243,12 @@ end end pretrivialization open pretrivialization -variables (F₁ E₁ F₂ E₂) [ring_hom_isometric σ] +variables (F₁ E₁ F₂ E₂) variables [Π x : B, topological_space (E₁ x)] [fiber_bundle F₁ E₁] [vector_bundle 𝕜₁ F₁ E₁] variables [Π x : B, topological_space (E₂ x)] [fiber_bundle F₂ E₂] [vector_bundle 𝕜₂ F₂ E₂] -variables [Π x, has_continuous_add (E₂ x)] [Π x, has_continuous_smul 𝕜₂ (E₂ x)] +variables [Π x, topological_add_group (E₂ x)] [Π x, has_continuous_smul 𝕜₂ (E₂ x)] + +include iσ /-- The continuous `σ`-semilinear maps between two topological vector bundles form a `vector_prebundle` (this is an auxiliary construction for the @@ -259,14 +275,25 @@ def _root_.bundle.continuous_linear_map.vector_prebundle : resetI, exact ⟨continuous_linear_map_coord_change σ e₁ e₁' e₂ e₂', continuous_on_continuous_linear_map_coord_change, - continuous_linear_map_coord_change_apply σ e₁ e₁' e₂ e₂'⟩ } } - -/-- Topology on the continuous `σ`-semilinear_maps between the respective fibers at a point of two -"normable" vector bundles over the same base. Here "normable" means that the bundles have fibers -modelled on normed spaces `F₁`, `F₂` respectively. The topology we put on the continuous -`σ`-semilinear_maps is the topology coming from the operator norm on maps from `F₁` to `F₂`. -/ -instance (x : B) : topological_space (bundle.continuous_linear_map σ F₁ E₁ F₂ E₂ x) := -(bundle.continuous_linear_map.vector_prebundle σ F₁ E₁ F₂ E₂).fiber_topology x + continuous_linear_map_coord_change_apply σ e₁ e₁' e₂ e₂'⟩ }, + total_space_mk_inducing := + begin + intros b, + dsimp [bundle.continuous_linear_map.topological_space, bundle.continuous_linear_map], + let L₁ : E₁ b ≃L[𝕜₁] F₁ := (trivialization_at F₁ E₁ b).continuous_linear_equiv_at 𝕜₁ b + (mem_base_set_trivialization_at _ _ _), + let L₂ : E₂ b ≃L[𝕜₂] F₂ := (trivialization_at F₂ E₂ b).continuous_linear_equiv_at 𝕜₂ b + (mem_base_set_trivialization_at _ _ _), + let φ : (E₁ b →SL[σ] E₂ b) ≃L[𝕜₂] (F₁ →SL[σ] F₂) := L₁.arrow_congrSL L₂, + have : inducing (λ x, (b, φ x)) := inducing_const_prod.mpr φ.to_homeomorph.inducing, + convert this, + ext f, + { refl }, + ext x, + dsimp [φ, pretrivialization.continuous_linear_map_apply], + rw [trivialization.linear_map_at_def_of_mem _ (mem_base_set_trivialization_at _ _ _)], + refl + end } /-- Topology on the total space of the continuous `σ`-semilinear_maps between two "normable" vector bundles over the same base. -/ @@ -313,7 +340,8 @@ rfl lemma trivialization.continuous_linear_map_apply (p : total_space (bundle.continuous_linear_map σ F₁ E₁ F₂ E₂)) : e₁.continuous_linear_map σ e₂ p = - ⟨p.1, (e₂.continuous_linear_map_at 𝕜₂ p.1).comp $ p.2.comp $ e₁.symmL 𝕜₁ p.1⟩ := + ⟨p.1, (e₂.continuous_linear_map_at 𝕜₂ p.1 : _ →L[𝕜₂] _).comp + (p.2.comp (e₁.symmL 𝕜₁ p.1 : F₁ →L[𝕜₁] E₁ p.1) : F₁ →SL[σ] E₂ p.1)⟩ := rfl omit he₁ he₂ From 5c4b3d41a84bd2a1d79c7d9265e58a891e71be89 Mon Sep 17 00:00:00 2001 From: damiano Date: Sun, 28 May 2023 16:55:33 +0000 Subject: [PATCH 1104/1291] chore(ring_theory/adjoin_root): modify `field` instance on `adjoin_root` (#19119) We copy the fields of `group_with_zero` instead of `field`. This speeds up the already ported `adjoin_root/AdjoinRoot` pair. [Zulip discussion](https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/!4.234339.20.20FieldTheory.2ESplittingField) --- src/ring_theory/adjoin_root.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/ring_theory/adjoin_root.lean b/src/ring_theory/adjoin_root.lean index f548e2fc730b2..efb922d502378 100644 --- a/src/ring_theory/adjoin_root.lean +++ b/src/ring_theory/adjoin_root.lean @@ -322,7 +322,7 @@ noncomputable instance field [fact (irreducible f)] : field (adjoin_root f) := qsmul_eq_mul' := λ a x, adjoin_root.induction_on _ x (λ p, by { rw [smul_mk, of, ring_hom.comp_apply, ← (mk f).map_mul, polynomial.rat_smul_eq_C_mul] }), ..adjoin_root.comm_ring f, - ..ideal.quotient.field (span {f} : ideal K[X]) } + ..ideal.quotient.group_with_zero (span {f} : ideal K[X]) } lemma coe_injective (h : degree f ≠ 0) : function.injective (coe : K → adjoin_root f) := have _ := adjoin_root.nontrivial f h, by exactI (of f).injective From 7bf0850d06ef8c83f234bfbe7dc5e90f422eaaa1 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sun, 28 May 2023 19:28:50 +0000 Subject: [PATCH 1105/1291] fix(scripts): Allow special character in file names (#11001) We don't have any files like this in mathlib, but it probably doesn't hurt to be safe here. --- scripts/mk_all.sh | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/scripts/mk_all.sh b/scripts/mk_all.sh index b9679e7211308..718163249628e 100755 --- a/scripts/mk_all.sh +++ b/scripts/mk_all.sh @@ -18,9 +18,9 @@ fi # remove an initial `./` # replace an initial `../test/` with just `.` (similarly for `roadmap`/`archive`/...) -# replace all `/` with `.` -# strip the `.lean` suffix -# prepend `import ` +# replace all `/` with `».«` +# replace the `.lean` suffix with `»` +# prepend `import «` find "$dir" -name \*.lean -not -name all.lean \ - | sed 's,^\./,,;s,^\.\./[^/]*,,;s,/,.,g;s,\.lean$,,;s,^,import ,' \ + | sed 's,^\./,,;s,^\.\./[^/]*,,;s,/,».«,g;s,\.lean$,»,;s,^,import «,' \ | sort >"$dir"/all.lean From 88a563b158f59f2983cfad685664da95502e8cdd Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sun, 28 May 2023 20:45:53 +0000 Subject: [PATCH 1106/1291] feat(analysis/normed_space/triv_sq_zero_ext): generalize some results to non-commutativity (#19049) --- .../normed_space/triv_sq_zero_ext.lean | 99 ++++++++++++++----- 1 file changed, 74 insertions(+), 25 deletions(-) diff --git a/src/analysis/normed_space/triv_sq_zero_ext.lean b/src/analysis/normed_space/triv_sq_zero_ext.lean index 53f5ced3084a8..a21d6f4182856 100644 --- a/src/analysis/normed_space/triv_sq_zero_ext.lean +++ b/src/analysis/normed_space/triv_sq_zero_ext.lean @@ -20,9 +20,14 @@ For now, this file contains results about `exp` for this type. * `triv_sq_zero_ext.exp_inr` ## TODO + * Actually define a sensible norm on `triv_sq_zero_ext R M`, so that we have access to lemmas like `exp_add`. -* Generalize some of these results to non-commutative `R`. +* Generalize more of these results to non-commutative `R`. In principle, under sufficient conditions + we should expect + `(exp 𝕜 x).snd = ∫ t in 0..1, exp 𝕜 (t • x.fst) • op (exp 𝕜 ((1 - t) • x.fst)) • x.snd` + ([Physics.SE](https://physics.stackexchange.com/a/41671/185147), and + https://link.springer.com/chapter/10.1007/978-3-540-44953-9_2). -/ @@ -35,24 +40,35 @@ namespace triv_sq_zero_ext section topology variables [topological_space R] [topological_space M] - -/-- If `exp R x.fst` converges to `e` then `exp R x` converges to `inl e + inr (e • x.snd)`. -/ -lemma has_sum_exp_series [field 𝕜] [char_zero 𝕜] [comm_ring R] +/-- If `exp R x.fst` converges to `e` then `(exp R x).fst` converges to `e`. -/ +lemma has_sum_fst_exp_series [field 𝕜] [ring R] [add_comm_group M] [algebra 𝕜 R] - [module R M] [module Rᵐᵒᵖ M] [is_central_scalar R M] - [module 𝕜 M] [is_scalar_tower 𝕜 R M] + [module R M] [module Rᵐᵒᵖ M] [smul_comm_class R Rᵐᵒᵖ M] + [module 𝕜 M] [is_scalar_tower 𝕜 R M] [is_scalar_tower 𝕜 Rᵐᵒᵖ M] [topological_ring R] [topological_add_group M] - [has_continuous_smul R M] - (x : tsze R M) {e : R} (h : has_sum (λ n, exp_series 𝕜 R n (λ _, x.fst)) e) : - has_sum (λ n, exp_series 𝕜 (tsze R M) n (λ _, x)) (inl e + inr (e • x.snd)) := + [has_continuous_smul R M] [has_continuous_smul Rᵐᵒᵖ M] + (x : tsze R M) + {e : R} (h : has_sum (λ n, exp_series 𝕜 R n (λ _, x.fst)) e) : + has_sum (λ n, fst (exp_series 𝕜 (tsze R M) n (λ _, x))) e := +by simpa [exp_series_apply_eq] using h + +/-- If `exp R x.fst` converges to `e` then `(exp R x).snd` converges to `e • x.snd`. -/ +lemma has_sum_snd_exp_series_of_smul_comm [field 𝕜] [char_zero 𝕜] [ring R] + [add_comm_group M] [algebra 𝕜 R] + [module R M] [module Rᵐᵒᵖ M] [smul_comm_class R Rᵐᵒᵖ M] + [module 𝕜 M] [is_scalar_tower 𝕜 R M] [is_scalar_tower 𝕜 Rᵐᵒᵖ M] + [topological_ring R] [topological_add_group M] + [has_continuous_smul R M] [has_continuous_smul Rᵐᵒᵖ M] + (x : tsze R M) (hx : mul_opposite.op x.fst • x.snd = x.fst • x.snd) + {e : R} (h : has_sum (λ n, exp_series 𝕜 R n (λ _, x.fst)) e) : + has_sum (λ n, snd (exp_series 𝕜 (tsze R M) n (λ _, x))) (e • x.snd) := begin simp_rw [exp_series_apply_eq] at *, conv { congr, funext, - rw [←inl_fst_add_inr_snd_eq (x ^ _), fst_pow, snd_pow, smul_add, ←inr_smul, - ←inl_smul, nsmul_eq_smul_cast 𝕜 n, smul_smul, inv_mul_eq_div, ←inv_div, ←smul_assoc], }, - refine (has_sum_inl M h).add (has_sum_inr M _), + rw [snd_smul, snd_pow_of_smul_comm _ _ hx, nsmul_eq_smul_cast 𝕜 n, smul_smul, inv_mul_eq_div, + ←inv_div, ←smul_assoc], }, apply has_sum.smul_const, rw [←has_sum_nat_add_iff' 1], swap, apply_instance, rw [finset.range_one, finset.sum_singleton, nat.cast_zero, div_zero, inv_zero, zero_smul, @@ -63,9 +79,53 @@ begin exact h, end +/-- If `exp R x.fst` converges to `e` then `exp R x` converges to `inl e + inr (e • x.snd)`. -/ +lemma has_sum_exp_series_of_smul_comm [field 𝕜] [char_zero 𝕜] [ring R] + [add_comm_group M] [algebra 𝕜 R] + [module R M] [module Rᵐᵒᵖ M] [smul_comm_class R Rᵐᵒᵖ M] + [module 𝕜 M] [is_scalar_tower 𝕜 R M] [is_scalar_tower 𝕜 Rᵐᵒᵖ M] + [topological_ring R] [topological_add_group M] + [has_continuous_smul R M] [has_continuous_smul Rᵐᵒᵖ M] + (x : tsze R M) (hx : mul_opposite.op x.fst • x.snd = x.fst • x.snd) + {e : R} (h : has_sum (λ n, exp_series 𝕜 R n (λ _, x.fst)) e) : + has_sum (λ n, exp_series 𝕜 (tsze R M) n (λ _, x)) (inl e + inr (e • x.snd)) := +by simpa only [inl_fst_add_inr_snd_eq] using + (has_sum_inl _ $ has_sum_fst_exp_series 𝕜 x h).add + (has_sum_inr _ $ has_sum_snd_exp_series_of_smul_comm 𝕜 x hx h) + end topology section normed_ring +variables [is_R_or_C 𝕜] [normed_ring R] [add_comm_group M] +variables [normed_algebra 𝕜 R] [module R M] [module Rᵐᵒᵖ M] [smul_comm_class R Rᵐᵒᵖ M] +variables [module 𝕜 M] [is_scalar_tower 𝕜 R M] [is_scalar_tower 𝕜 Rᵐᵒᵖ M] +variables [topological_space M] [topological_ring R] +variables [topological_add_group M] [has_continuous_smul R M] [has_continuous_smul Rᵐᵒᵖ M] +variables [complete_space R] [t2_space R] [t2_space M] + +lemma exp_def_of_smul_comm (x : tsze R M) (hx : mul_opposite.op x.fst • x.snd = x.fst • x.snd) : + exp 𝕜 x = inl (exp 𝕜 x.fst) + inr (exp 𝕜 x.fst • x.snd) := +begin + simp_rw [exp, formal_multilinear_series.sum], + refine (has_sum_exp_series_of_smul_comm 𝕜 x hx _).tsum_eq, + exact exp_series_has_sum_exp _, +end + +@[simp] lemma exp_inl (x : R) : exp 𝕜 (inl x : tsze R M) = inl (exp 𝕜 x) := +begin + rw [exp_def_of_smul_comm, snd_inl, fst_inl, smul_zero, inr_zero, add_zero], + { rw [snd_inl, fst_inl, smul_zero, smul_zero] } +end + +@[simp] lemma exp_inr (m : M) : exp 𝕜 (inr m : tsze R M) = 1 + inr m := +begin + rw [exp_def_of_smul_comm, snd_inr, fst_inr, exp_zero, one_smul, inl_one], + { rw [snd_inr, fst_inr, mul_opposite.op_zero, zero_smul, zero_smul] } +end + +end normed_ring + +section normed_comm_ring variables [is_R_or_C 𝕜] [normed_comm_ring R] [add_comm_group M] variables [normed_algebra 𝕜 R] [module R M] [module Rᵐᵒᵖ M] [is_central_scalar R M] variables [module 𝕜 M] [is_scalar_tower 𝕜 R M] @@ -74,11 +134,7 @@ variables [topological_add_group M] [has_continuous_smul R M] variables [complete_space R] [t2_space R] [t2_space M] lemma exp_def (x : tsze R M) : exp 𝕜 x = inl (exp 𝕜 x.fst) + inr (exp 𝕜 x.fst • x.snd) := -begin - simp_rw [exp, formal_multilinear_series.sum], - refine (has_sum_exp_series 𝕜 x _).tsum_eq, - exact exp_series_has_sum_exp _, -end +exp_def_of_smul_comm 𝕜 x (op_smul_eq_smul _ _) @[simp] lemma fst_exp (x : tsze R M) : fst (exp 𝕜 x) = exp 𝕜 x.fst := by rw [exp_def, fst_add, fst_inl, fst_inr, add_zero] @@ -86,20 +142,13 @@ by rw [exp_def, fst_add, fst_inl, fst_inr, add_zero] @[simp] lemma snd_exp (x : tsze R M) : snd (exp 𝕜 x) = exp 𝕜 x.fst • x.snd := by rw [exp_def, snd_add, snd_inl, snd_inr, zero_add] -@[simp] lemma exp_inl (x : R) : exp 𝕜 (inl x : tsze R M) = inl (exp 𝕜 x) := -by rw [exp_def, fst_inl, snd_inl, smul_zero, inr_zero, add_zero] - -@[simp] lemma exp_inr (m : M) : exp 𝕜 (inr m : tsze R M) = 1 + inr m := -by rw [exp_def, fst_inr, exp_zero, snd_inr, one_smul, inl_one] - /-- Polar form of trivial-square-zero extension. -/ lemma eq_smul_exp_of_invertible (x : tsze R M) [invertible x.fst] : x = x.fst • exp 𝕜 (⅟x.fst • inr x.snd) := by rw [←inr_smul, exp_inr, smul_add, ←inl_one, ←inl_smul, ←inr_smul, smul_eq_mul, mul_one, smul_smul, mul_inv_of_self, one_smul, inl_fst_add_inr_snd_eq] -end normed_ring - +end normed_comm_ring section normed_field variables [is_R_or_C 𝕜] [normed_field R] [add_comm_group M] From e3e30a5c801cd7ce00eeb3644200016669f9ed3e Mon Sep 17 00:00:00 2001 From: Moritz Doll Date: Mon, 29 May 2023 05:08:40 +0000 Subject: [PATCH 1107/1291] feat(analysis/schwartz_space): partial derivatives (#18755) --- src/analysis/schwartz_space.lean | 60 ++++++++++++++++++++++++++++++++ 1 file changed, 60 insertions(+) diff --git a/src/analysis/schwartz_space.lean b/src/analysis/schwartz_space.lean index 42b870cd55355..b7971ba0335e9 100644 --- a/src/analysis/schwartz_space.lean +++ b/src/analysis/schwartz_space.lean @@ -575,6 +575,27 @@ def mk_clm [ring_hom_isometric σ] (A : (D → E) → (F → G)) end clm +section eval_clm + +variables [normed_field 𝕜] [normed_space 𝕜 F] [smul_comm_class ℝ 𝕜 F] + +/-- The map applying a vector to Hom-valued Schwartz function as a continuous linear map. -/ +@[protected] def eval_clm (m : E) : 𝓢(E, E →L[ℝ] F) →L[𝕜] 𝓢(E, F) := +mk_clm (λ f x, f x m) + (λ _ _ _, rfl) (λ _ _ _, rfl) (λ f, cont_diff.clm_apply f.2 cont_diff_const) + (begin + rintro ⟨k, n⟩, + use [{(k, n)}, ‖m‖, norm_nonneg _], + intros f x, + refine le_trans (mul_le_mul_of_nonneg_left (norm_iterated_fderiv_clm_apply_const f.2 le_top) + (by positivity)) _, + rw [← mul_assoc, ← mul_comm (‖m‖), mul_assoc], + refine mul_le_mul_of_nonneg_left _ (norm_nonneg _), + simp only [finset.sup_singleton, schwartz_seminorm_family_apply, le_seminorm], + end) + +end eval_clm + section derivatives /-! ### Derivatives of Schwartz functions -/ @@ -608,6 +629,45 @@ mk_clm (λ f, deriv f) @[simp] lemma deriv_clm_apply (f : 𝓢(ℝ, F)) (x : ℝ) : deriv_clm 𝕜 f x = deriv f x := rfl +/-- The partial derivative (or directional derivative) in the direction `m : E` as a +continuous linear map on Schwartz space. -/ +def pderiv_clm (m : E) : 𝓢(E, F) →L[𝕜] 𝓢(E, F) := (eval_clm m).comp (fderiv_clm 𝕜) + +@[simp] +lemma pderiv_clm_apply (m : E) (f : 𝓢(E, F)) (x : E) : pderiv_clm 𝕜 m f x = fderiv ℝ f x m := rfl + +/-- The iterated partial derivative (or directional derivative) as a continuous linear map on +Schwartz space. -/ +def iterated_pderiv {n : ℕ} : (fin n → E) → 𝓢(E, F) →L[𝕜] 𝓢(E, F) := +nat.rec_on n + (λ x, continuous_linear_map.id 𝕜 _) + (λ n rec x, (pderiv_clm 𝕜 (x 0)).comp (rec (fin.tail x))) + +@[simp] lemma iterated_pderiv_zero (m : fin 0 → E) (f : 𝓢(E, F)): + iterated_pderiv 𝕜 m f = f := rfl + +@[simp] lemma iterated_pderiv_one (m : fin 1 → E) (f : 𝓢(E, F)) : + iterated_pderiv 𝕜 m f = pderiv_clm 𝕜 (m 0) f := rfl + +lemma iterated_pderiv_succ_left {n : ℕ} (m : fin (n + 1) → E) (f : 𝓢(E, F)) : + iterated_pderiv 𝕜 m f = pderiv_clm 𝕜 (m 0) (iterated_pderiv 𝕜 (fin.tail m) f) := rfl + +lemma iterated_pderiv_succ_right {n : ℕ} (m : fin (n + 1) → E) (f : 𝓢(E, F)) : + iterated_pderiv 𝕜 m f = + iterated_pderiv 𝕜 (fin.init m) (pderiv_clm 𝕜 (m (fin.last n)) f) := +begin + induction n with n IH, + { rw [iterated_pderiv_zero, iterated_pderiv_one], + refl }, + -- The proof is `∂^{n + 2} = ∂ ∂^{n + 1} = ∂ ∂^n ∂ = ∂^{n+1} ∂` + have hmzero : fin.init m 0 = m 0 := by simp only [fin.init_def, fin.cast_succ_zero], + have hmtail : fin.tail m (fin.last n) = m (fin.last n.succ) := + by simp only [fin.tail_def, fin.succ_last], + simp only [iterated_pderiv_succ_left, IH (fin.tail m), hmzero, hmtail, fin.tail_init_eq_init_tail] +end + +-- Todo: `iterated_pderiv 𝕜 m f x = iterated_fderiv ℝ f x m` + end derivatives section bounded_continuous_function From a84108922d25a66af47d2535e4ce6de188f85192 Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Mon, 29 May 2023 08:34:45 +0000 Subject: [PATCH 1108/1291] feat (number_theory/zeta_function): define zeta (#19027) This PR defines the Riemann zeta function (as a function on the whole complex plane). --- src/number_theory/zeta_function.lean | 351 +++++++++++++++++++++++++++ 1 file changed, 351 insertions(+) create mode 100644 src/number_theory/zeta_function.lean diff --git a/src/number_theory/zeta_function.lean b/src/number_theory/zeta_function.lean new file mode 100644 index 0000000000000..20b3b7a2b3991 --- /dev/null +++ b/src/number_theory/zeta_function.lean @@ -0,0 +1,351 @@ +/- +Copyright (c) 2023 David Loeffler. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: David Loeffler +-/ +import analysis.mellin_transform +import number_theory.modular_forms.jacobi_theta + +/-! +# Definition of the Riemann zeta function + +## Main definitions: + +* `riemann_zeta`: the Riemann zeta function `ζ : ℂ → ℂ`. +* `riemann_completed_zeta`: the completed zeta function `Λ : ℂ → ℂ`, which satisfies + `Λ(s) = π ^ (-s / 2) Γ(s / 2) ζ(s)` (away from the poles of `Γ(s / 2)`). +* `riemann_completed_zeta₀`: the entire function `Λ₀` satisfying + `Λ₀(s) = Λ(s) + 1 / (s - 1) - 1 / s` wherever the RHS is defined. + +Note that mathematically `ζ(s)` is undefined at `s = 1`, while `Λ(s)` is undefined at both `s = 0` +and `s = 1`. Our construction assigns some values at these points (which are not arbitrary, but +I haven't checked exactly what they are). + +## Main results: + +* `differentiable_completed_zeta₀` : the function `Λ₀(s)` is entire. +* `differentiable_at_completed_zeta`: the function `Λ(s)` is differentiable away from `s = 0` and + `s = 1`. + +## Outline of proofs: + +We define two related functions on the reals, `zeta_kernel₁` and `zeta_kernel₂`. The first is +`(Θ (t * I) - 1) / 2`, where `θ` is Jacobi's theta function; its Mellin transform is exactly the +completed zeta function. The second is obtained by subtracting a linear combination of powers on +the interval `Ioc 0 1` to give a function with exponential decay at both `0` and `∞`. We then define +`riemann_completed_zeta₀` as the Mellin transform of the second zeta kernel, and define +`riemann_completed_zeta` and `riemann_zeta` from this. +-/ + +open measure_theory set filter asymptotics topological_space real asymptotics +open complex (hiding exp norm_eq_abs abs_of_nonneg abs_two continuous_exp) + +open_locale topology real + +noncomputable theory + +/-! +## Definition of the Riemann zeta function and related functions +-/ + +/-- Function whose Mellin transform is `π ^ (-s) * Γ(s) * zeta (2 * s)`, for `1 / 2 < Re s`. -/ +def zeta_kernel₁ (t : ℝ) : ℂ := ∑' (n : ℕ), rexp (-π * t * (n + 1) ^ 2) + +/-- Modified zeta kernel, whose Mellin transform is entire. --/ +def zeta_kernel₂ : ℝ → ℂ := zeta_kernel₁ + indicator (Ioc 0 1) (λ t, (1 - 1 / sqrt t) / 2) + +/-- The completed Riemann zeta function with its poles removed, `Λ(s) + 1 / s - 1 / (s - 1)`. -/ +def riemann_completed_zeta₀ (s : ℂ) : ℂ := mellin zeta_kernel₂ (s / 2) + +/-- The completed Riemann zeta function, `Λ(s)`, which satisfies +`Λ(s) = π ^ (-s / 2) Γ(s / 2) ζ(s)` (up to a minor correction at `s = 0`). -/ +def riemann_completed_zeta (s : ℂ) : ℂ := riemann_completed_zeta₀ s - 1 / s + 1 / (s - 1) + +/-- The Riemann zeta function `ζ(s)`. We set this to be irreducible to hide messy implementation +details. -/ +@[irreducible] def riemann_zeta := function.update + (λ s : ℂ, ↑π ^ (s / 2) * riemann_completed_zeta s / Gamma (s / 2)) 0 (-1 / 2) + +/- Note the next lemma is true by definition; what's hard is to show that with this definition, `ζ` +is continuous (and indeed analytic) at 0. + +TODO: Prove this -- this will be in a future PR. -/ +/-- We have `ζ(0) = -1 / 2`. -/ +lemma riemann_zeta_zero : riemann_zeta 0 = -1 / 2 := +begin + unfold riemann_zeta, + exact function.update_same _ _ _ +end + +/-! +## First properties of the zeta kernels +-/ + +/-- The sum defining `zeta_kernel₁` is convergent. -/ +lemma summable_exp_neg_pi_mul_nat_sq {t : ℝ} (ht : 0 < t) : + summable (λ n : ℕ, rexp (-π * t * (n + 1) ^ 2)) := +begin + have : 0 < (↑t * I).im, by rwa [of_real_mul_im, I_im, mul_one], + convert summable_norm_iff.mpr (has_sum_nat_jacobi_theta this).summable, + ext1 n, + rw [complex.norm_eq_abs, complex.abs_exp], + rw show ↑π * I * (↑n + 1) ^ 2 * (↑t * I) = ↑(π * t * (n + 1) ^ 2) * I ^ 2, by { push_cast, ring }, + rw [I_sq, mul_neg_one, ←of_real_neg, of_real_re, neg_mul, neg_mul], +end + +/-- Relate `zeta_kernel₁` to the Jacobi theta function on `ℍ`. (We don't use this as the definition +of `zeta_kernel₁`, since the sum over `ℕ` rather than `ℤ` is more convenient for relating zeta to +the Dirichlet series for `re s > 1`.) -/ +lemma zeta_kernel₁_eq_jacobi_theta {t : ℝ} (ht : 0 < t) : + zeta_kernel₁ t = (jacobi_theta (t * I) - 1) / 2 := +begin + rw [jacobi_theta_eq_tsum_nat ((mul_I_im t).symm ▸ ht : 0 < (↑t * I).im), add_comm, add_sub_cancel, + mul_div_cancel_left _ (two_ne_zero' ℂ), zeta_kernel₁], + congr' 1 with n : 1, + push_cast, + rw [(by ring : ↑π * I * (n + 1) ^ 2 * (t * I) = I ^ 2 * π * t * (n + 1) ^ 2), I_sq, neg_one_mul], +end + +/-- Continuity of `zeta_kernel₁`. -/ +lemma continuous_at_zeta_kernel₁ {t : ℝ} (ht : 0 < t) : continuous_at zeta_kernel₁ t := +begin + have : continuous_at (λ u : ℝ, (jacobi_theta (u * I) - 1) / 2) t, + { refine (continuous_at.sub _ continuous_at_const).div_const _, + refine (continuous_at_jacobi_theta _).comp (continuous_at.mul _ continuous_at_const), + { rwa [mul_I_im, of_real_re] }, + { exact continuous_of_real.continuous_at } }, + refine this.congr (eventually_of_mem (Ioi_mem_nhds ht) (λ u hu, _)), + rw zeta_kernel₁_eq_jacobi_theta hu, +end + +/-- Local integrability of `zeta_kernel₁`. -/ +lemma locally_integrable_zeta_kernel₁ : locally_integrable_on zeta_kernel₁ (Ioi 0) := +(continuous_at.continuous_on $ λ t ht, continuous_at_zeta_kernel₁ ht).locally_integrable_on + measurable_set_Ioi + +/-- Local integrability of `zeta_kernel₂`. -/ +lemma locally_integrable_zeta_kernel₂ : locally_integrable_on zeta_kernel₂ (Ioi 0) := +begin + refine (locally_integrable_on_iff (or.inr is_open_Ioi)).mpr (λ k hk hk', integrable.add _ _), + { refine continuous_on.integrable_on_compact hk' _, + exact continuous_at.continuous_on (λ x hx, continuous_at_zeta_kernel₁ (hk hx)) }, + { refine (integrable_indicator_iff measurable_set_Ioc).mpr _, + rw [integrable_on, measure.restrict_restrict, ←integrable_on], + swap, { exact measurable_set_Ioc }, + apply continuous_on.integrable_on_compact, + { convert (is_compact_Icc : is_compact $ Icc 0 1).inter hk' using 1, + exact set.ext (λ t, ⟨λ h, ⟨Ioc_subset_Icc_self h.1, h.2⟩, λ h, ⟨⟨hk h.2, h.1.2⟩, h.2⟩⟩) }, + { refine continuous_on.mono _ ((inter_subset_right _ _).trans hk), + refine (continuous_on_const.sub _).div_const _, + refine continuous_on.div continuous_on_const _ (λ x hx, _), + { exact (continuous_of_real.comp continuous_sqrt).continuous_on }, + exact of_real_ne_zero.mpr (sqrt_ne_zero'.mpr hx) } } +end + +/-- Functional equation for `zeta_kernel₂`. -/ +lemma zeta_kernel₂_one_div {t : ℝ} (ht : 0 < t) : + zeta_kernel₂ (1 / t) = sqrt t * zeta_kernel₂ t := +begin + have aux : ∀ {u : ℝ} (hu : 1 < u), zeta_kernel₂ (1 / u) = sqrt u * zeta_kernel₂ u, + { intros u hu, + simp_rw [zeta_kernel₂, pi.add_apply], + rw [indicator_of_mem, indicator_of_not_mem (not_mem_Ioc_of_gt hu), add_zero], + swap, { exact ⟨one_div_pos.mpr (zero_lt_one.trans hu), (one_div u).symm ▸ (inv_le_one hu.le)⟩ }, + rw [zeta_kernel₁_eq_jacobi_theta (one_div_pos.mpr $ zero_lt_one.trans hu), + zeta_kernel₁_eq_jacobi_theta (zero_lt_one.trans hu), ←add_div, ←mul_div_assoc, add_sub, + sub_add_cancel, sqrt_div zero_le_one, sqrt_one, one_div (sqrt _), of_real_inv, + ←one_div, one_div_one_div, mul_sub, mul_one], + congr' 2, + let τ : upper_half_plane := ⟨u * I, (mul_I_im u).symm ▸ (zero_lt_one.trans hu)⟩, + convert jacobi_theta_S_smul τ using 2, + { rw [upper_half_plane.modular_S_smul, upper_half_plane.coe_mk, subtype.coe_mk, ←neg_inv, + mul_inv, inv_I, mul_neg, neg_neg, one_div, of_real_inv], }, + { rw [subtype.coe_mk, mul_comm, mul_assoc, mul_neg, I_mul_I, neg_neg, mul_one, + sqrt_eq_rpow, of_real_cpow (zero_lt_one.trans hu).le], + push_cast } }, + rcases lt_trichotomy 1 t with h | rfl | h, + { exact aux h }, + { simp only [div_self, ne.def, one_ne_zero, not_false_iff, sqrt_one, of_real_one, one_mul], }, + { have := aux (show 1 < 1 / t, by rwa [lt_one_div (zero_lt_one' ℝ) ht, div_one]), + rw one_div_one_div at this, + rw [this, ←mul_assoc, ←of_real_mul, ←sqrt_mul ht.le, mul_one_div_cancel ht.ne', sqrt_one, + of_real_one, one_mul] }, +end + +/-! +## Bounds for zeta kernels + +We now establish asymptotic bounds for the zeta kernels as `t → ∞` and `t → 0`, and use these to +show holomorphy of their Mellin transforms (for `1 / 2 < re s` for `zeta_kernel₁`, and all `s` for +`zeta_kernel₂`). -/ + +/-- Bound for `zeta_kernel₁` for large `t`. -/ +lemma is_O_at_top_zeta_kernel₁ : is_O at_top zeta_kernel₁ (λ t, exp (-π * t)) := +begin + have h := (is_O_at_im_infty_jacobi_theta_sub_one).const_mul_left (1 / 2), + simp_rw [(mul_comm (1 / 2 : ℂ) _), mul_one_div] at h, + have h' : tendsto (λ t : ℝ, ↑t * I) at_top (comap im at_top), + { rw tendsto_comap_iff, + convert tendsto_id, + ext1 t, + rw [function.comp_app, mul_I_im, of_real_re, id.def] }, + convert ((h.norm_left.comp_tendsto h').congr' (eventually_of_mem (Ioi_mem_at_top 0) (λ t ht, _)) + (eventually_of_mem (Ioi_mem_at_top 0) (λ t ht, _))).of_norm_left, + { rw [function.comp_app, ←zeta_kernel₁_eq_jacobi_theta ht] }, + { rw [function.comp_app, mul_I_im, of_real_re] } +end + +/-- Bound for `zeta_kernel₂` for large `t`. -/ +lemma is_O_at_top_zeta_kernel₂ : is_O at_top zeta_kernel₂ (λ t, exp (-π * t)) := +begin + refine (eventually_eq_of_mem (Ioi_mem_at_top (1 : ℝ)) (λ t ht, _)).trans_is_O + is_O_at_top_zeta_kernel₁, + rw [zeta_kernel₂, pi.add_apply, indicator_of_not_mem (not_mem_Ioc_of_gt ht), add_zero], +end + +/-- Precise but awkward-to-use bound for `zeta_kernel₂` for `t → 0`. -/ +lemma is_O_zero_zeta_kernel₂ : is_O (𝓝[>] 0) zeta_kernel₂ (λ t, exp (-π / t) / sqrt t) := +begin + have h1 := (is_O_at_top_zeta_kernel₂).comp_tendsto tendsto_inv_zero_at_top, + simp_rw ←one_div at h1, + have h2 : (zeta_kernel₂ ∘ has_div.div 1) =ᶠ[𝓝[>] 0] λ t, sqrt t * zeta_kernel₂ t, + from eventually_of_mem self_mem_nhds_within (λ t ht, by simp_rw ←zeta_kernel₂_one_div ht), + have h3 := (h1.congr' h2 (eventually_eq.refl _ _)), + have h4 := h3.mul (is_O_refl (λ t : ℝ, 1 / (sqrt t : ℂ)) (𝓝[>] 0)).norm_right, + refine h4.congr' _ _, + { refine eventually_of_mem self_mem_nhds_within (λ x hx, _), + simp_rw [←mul_assoc], + rw [mul_comm, ←mul_assoc, one_div_mul_cancel, one_mul], + exact of_real_ne_zero.mpr ((sqrt_ne_zero $ le_of_lt hx).mpr (ne_of_gt hx)) }, + { refine eventually_of_mem self_mem_nhds_within (λ x hx, _), + dsimp only, + rw [function.comp_app, mul_one_div, one_div (↑(sqrt _)), ←of_real_inv, is_R_or_C.norm_of_real, + abs_inv, abs_of_nonneg (sqrt_nonneg _), ←div_eq_mul_inv] }, +end + +/-- Weaker but more usable bound for `zeta_kernel₂` for `t → 0`. -/ +lemma is_O_zero_zeta_kernel₂_rpow (a : ℝ) : is_O (𝓝[>] 0) zeta_kernel₂ (λ t, t ^ a) := +begin + have aux1 : is_O at_top (λ t, exp (-π * t)) (λ t, t ^ (-a - 1 / 2)), + from (is_o_exp_neg_mul_rpow_at_top pi_pos _).is_O, + have aux2 : is_O at_top (λ t, exp (-π * t) * sqrt t) (λ t, t ^ (-a)), + { refine (aux1.mul (is_O_refl sqrt _)).congr' (eventually_eq.refl _ _) _, + refine (eventually_gt_at_top 0).mp (eventually_of_forall (λ t ht, _)), + simp_rw [sqrt_eq_rpow, ←rpow_add ht, sub_add_cancel] }, + refine is_O_zero_zeta_kernel₂.trans ((aux2.comp_tendsto tendsto_inv_zero_at_top).congr' _ _), + { refine eventually_of_mem self_mem_nhds_within (λ x hx, _), + simp_rw [function.comp_app, sqrt_inv, ←div_eq_mul_inv] }, + { refine eventually_of_mem self_mem_nhds_within (λ x hx, _), + simp_rw [function.comp_app, inv_rpow (le_of_lt hx), rpow_neg (le_of_lt hx), inv_inv] } +end + +/-- Bound for `zeta_kernel₁` for `t → 0`. -/ +lemma is_O_zero_zeta_kernel₁ : is_O (𝓝[>] 0) zeta_kernel₁ (λ t, t ^ (-(1 / 2) : ℝ)) := +begin + have : zeta_kernel₁ =ᶠ[𝓝[>] 0] zeta_kernel₂ + (λ t, (1 / sqrt t - 1) / 2), + { refine eventually_eq_of_mem (Ioc_mem_nhds_within_Ioi $ left_mem_Ico.mpr zero_lt_one) (λ t h, _), + rw [pi.add_apply, zeta_kernel₂, pi.add_apply, indicator_of_mem h], + ring }, + refine ((is_O_zero_zeta_kernel₂_rpow _).add _).congr' this.symm (eventually_eq.refl _ _), + simp_rw sub_div, + apply is_O.sub, + { apply is_O.of_norm_left, + simp_rw [norm_div, norm_one, div_eq_mul_inv, one_mul, mul_comm _ (‖(2 : ℂ)‖)⁻¹], + refine ((is_O_refl _ _).congr' (eventually_eq.refl _ _) + (eventually_eq_of_mem self_mem_nhds_within (λ x hx, _))).const_mul_left _, + rw [is_R_or_C.norm_of_real, abs_of_nonneg (sqrt_nonneg _)], + simp_rw [sqrt_eq_rpow, rpow_neg (le_of_lt hx), one_div] }, + { refine is_O_iff.mpr ⟨‖(1 / 2 : ℂ)‖, _⟩, + refine eventually_of_mem (Ioc_mem_nhds_within_Ioi $ left_mem_Ico.mpr zero_lt_one) (λ t ht, _), + refine le_mul_of_one_le_right (norm_nonneg _) _, + rw [norm_of_nonneg (rpow_nonneg_of_nonneg ht.1.le _), rpow_neg ht.1.le], + exact one_le_inv (rpow_pos_of_pos ht.1 _) (rpow_le_one ht.1.le ht.2 one_half_pos.le) } +end + +/-! +## Differentiability of the completed zeta function +-/ + +/-- The Mellin transform of the first zeta kernel is holomorphic for `1 / 2 < re s`. -/ +lemma differentiable_at_mellin_zeta_kernel₁ {s : ℂ} (hs : 1 / 2 < s.re) : + differentiable_at ℂ (mellin zeta_kernel₁) s := +mellin_differentiable_at_of_is_O_rpow_exp pi_pos locally_integrable_zeta_kernel₁ + is_O_at_top_zeta_kernel₁ is_O_zero_zeta_kernel₁ hs + +/-- The Mellin transform of the second zeta kernel is entire. -/ +lemma differentiable_mellin_zeta_kernel₂ : differentiable ℂ (mellin zeta_kernel₂) := +λ s, mellin_differentiable_at_of_is_O_rpow_exp pi_pos locally_integrable_zeta_kernel₂ + is_O_at_top_zeta_kernel₂ (is_O_zero_zeta_kernel₂_rpow _) ((sub_lt_self_iff _).mpr zero_lt_one) + +/-- The modified completed Riemann zeta function `Λ(s) + 1 / s - 1 / (s - 1)` is entire. -/ +theorem differentiable_completed_zeta₀ : differentiable ℂ riemann_completed_zeta₀ := +differentiable_mellin_zeta_kernel₂.comp (differentiable.div_const differentiable_id 2) + +/-- The completed Riemann zeta function `Λ(s)` is differentiable away from `s = 0` and `s = 1` +(where it has simple poles). -/ +theorem differentiable_at_completed_zeta {s : ℂ} (hs : s ≠ 0) (hs' : s ≠ 1) : + differentiable_at ℂ riemann_completed_zeta s := +begin + refine (differentiable_completed_zeta₀.differentiable_at.sub _).add _, + { exact (differentiable.differentiable_at (differentiable_const _)).div differentiable_at_id hs }, + { refine ((differentiable_const _).differentiable_at).div _ (sub_ne_zero.mpr hs'), + exact differentiable_at_id.sub (differentiable_at_const _) }, +end + +/-! +## Relating the Mellin transforms of the two zeta kernels +-/ + +lemma has_mellin_one_div_sqrt_Ioc {s : ℂ} (hs : 1 / 2 < re s) : + has_mellin (indicator (Ioc 0 1) (λ t, 1 / ↑(sqrt t) : ℝ → ℂ)) s (1 / (s - 1 / 2)) := +begin + have h1 : eq_on (λ t, 1 / ↑(sqrt t) : ℝ → ℂ) (λ t, ↑t ^ (-1 / 2 : ℂ)) (Ioc 0 1), + { intros t ht, + simp_rw [neg_div, cpow_neg, ←one_div, sqrt_eq_rpow, of_real_cpow ht.1.le], + push_cast }, + simp_rw [indicator_congr h1, (by ring : s - 1/2 = s + (-1) / 2)], + convert has_mellin_cpow_Ioc (-1 / 2) _, + rwa [(by push_cast : (-1 / 2 : ℂ) = (-1 / 2 : ℝ)), of_real_re, neg_div, ←sub_eq_add_neg, sub_pos] +end + +/-- Evaluate the Mellin transform of the "fudge factor" in `zeta_kernel₂` -/ +lemma has_mellin_one_div_sqrt_sub_one_div_two_Ioc {s : ℂ} (hs : 1 / 2 < s.re) : + has_mellin ((Ioc 0 1).indicator (λ t, (1 - 1 / (↑(sqrt t) : ℂ)) / 2)) s + (1 / (2 * s) - 1 / (2 * s - 1)) := +begin + have step1 : has_mellin (indicator (Ioc 0 1) (λ t, 1 - 1 / ↑(sqrt t) : ℝ → ℂ)) s + (1 / s - 1 / (s - 1 / 2)), + { have a := has_mellin_one_Ioc (one_half_pos.trans hs), + have b := has_mellin_one_div_sqrt_Ioc hs, + simpa only [a.2, b.2, ←indicator_sub] using has_mellin_sub a.1 b.1 }, + -- todo: implement something like "indicator.const_div" (blocked by the port for now) + rw (show (Ioc 0 1).indicator (λ t, (1 - 1 / (↑(sqrt t) : ℂ)) / 2) = + λ t, ((Ioc 0 1).indicator (λ t, (1 - 1 / (↑(sqrt t) : ℂ))) t) / 2, + by { ext1 t, simp_rw [div_eq_inv_mul, indicator_mul_right] }), + simp_rw [has_mellin, mellin_div_const, step1.2, sub_div, div_div], + refine ⟨step1.1.div_const _, _⟩, + rw [mul_comm, sub_mul, div_mul_cancel _ (two_ne_zero' ℂ), mul_comm s 2], +end + +lemma mellin_zeta_kernel₂_eq_of_lt_re {s : ℂ} (hs : 1 / 2 < s.re) : + mellin zeta_kernel₂ s = mellin zeta_kernel₁ s + 1 / (2 * s) - 1 / (2 * s - 1) := +begin + have h := mellin_convergent_of_is_O_rpow_exp pi_pos locally_integrable_zeta_kernel₁ + is_O_at_top_zeta_kernel₁ is_O_zero_zeta_kernel₁ hs, + have h' := has_mellin_one_div_sqrt_sub_one_div_two_Ioc hs, + simp_rw [zeta_kernel₂, pi.add_def, add_sub_assoc, (has_mellin_add h h'.1).2, h'.2], +end + +lemma completed_zeta_eq_mellin_of_one_lt_re {s : ℂ} (hs : 1 < re s) : + riemann_completed_zeta s = mellin zeta_kernel₁ (s / 2) := +begin + have : 1 / 2 < (s / 2).re, + { rw (show s / 2 = ↑(2⁻¹ : ℝ) * s, by { push_cast, rw mul_comm, refl }), + rwa [of_real_mul_re, ←div_eq_inv_mul, div_lt_div_right (zero_lt_two' ℝ)] }, + rw [riemann_completed_zeta, riemann_completed_zeta₀, mellin_zeta_kernel₂_eq_of_lt_re this, + sub_add, sub_sub, ←add_sub], + conv_rhs { rw ←add_zero (mellin zeta_kernel₁ $ s / 2) }, + congr' 1, + rw mul_div_cancel' _ (two_ne_zero' ℂ), + abel +end From 729d23f9e1640e1687141be89b106d3c8f9d10c0 Mon Sep 17 00:00:00 2001 From: "Filippo A. E. Nuccio" Date: Mon, 29 May 2023 16:53:40 +0000 Subject: [PATCH 1109/1291] doc(topology/homotopy/H-spaces): add doc about H-spaces. (#19075) Added doc about `H-spaces` in topology/homotopy/H-spaces. --- src/topology/homotopy/H_spaces.lean | 56 +++++++++++++++++++++++------ 1 file changed, 45 insertions(+), 11 deletions(-) diff --git a/src/topology/homotopy/H_spaces.lean b/src/topology/homotopy/H_spaces.lean index d39b2b8a4b8c0..722ad7fd54618 100644 --- a/src/topology/homotopy/H_spaces.lean +++ b/src/topology/homotopy/H_spaces.lean @@ -10,9 +10,34 @@ import topology.homotopy.path # H-spaces This file defines H-spaces mainly following the approach proposed by Serre in his paper -*Homologie singulière des espaces fibrés*. The main results are the H-space `instance` on every -topological group, and the H-space structure on the loop space (based at `x : X` of any topological -space `X`, for which we introduce the notation `Ω_[x]`. +*Homologie singulière des espaces fibrés*. The idea beaneath `H-spaces` is that they are topological +spaces with a binary operation `⋀ : X → X → X` that is a homotopic-theoretic weakening of an +operation what would make `X` into a topological monoid. In particular, there exists a "neutral +element" `e : X` such that `λ x, e ⋀ x` and `λ x, x ⋀ e` are homotopic to the identity on `X`, see +[the Wikipedia page of H-spaces](https://en.wikipedia.org/wiki/H-space). + +Some notable properties of `H-spaces` are +* Their fundamental group is always abelian (by the same argument for topological groups); +* Their cohomology ring comes equipped with a structure of a Hopf-algebra; +* The loop space based at every `x : X` carries a structure of an `H-spaces`. + +## Main Results + +* Every topological group `G` is an `H-space` using its operation `* : G → G → G` (this is already +true if `G` has an instance of a `mul_one_class` and `continuous_mul`); +* Given two `H-spaces` `X` and `Y`, their product is again an `H`-space. We show in an example that +starting with two topological groups `G, G'`, the `H`-space structure on `G × G'` is definitionally +equal to the product of `H-space` structures on `G` and `G'`. +* The loop space based at every `x : X` carries a structure of an `H-spaces`. + +## To Do +* Prove that for every `normed_add_torsor Z` and every `z : Z`, the operation +`λ x y, midpoint x y` defines a `H-space` structure with `z` as a "neutral element". +* Prove that `S^0`, `S^1`, `S^3` and `S^7` are the unique spheres that are `H-spaces`, where the +first three inherit the structure because they are topological groups (they are Lie groups, +actually), isomorphic to the invertible elements in `ℤ`, in `ℂ` and in the quaternion; and the +fourth from the fact that `S^7` coincides with the octonions of norm 1 (it is not a group, in +particular, only has an instance of `mul_one_class`). ## References @@ -91,25 +116,34 @@ instance H_space.prod (X : Type u) (Y : Type v) [topological_space X] [topologic namespace topological_group -@[priority 600, to_additive] instance H_space (G : Type u) [topological_space G] [group G] - [topological_group G] : H_space G := +/-- The definition `to_H_space` is not an instance because its `@additive` version would +lead to a diamond since a topological field would inherit two `H_space` structures, one from the +`mul_one_class` and one from the `add_zero_class`. In the case of a group, we make +`topological_group.H_space` an instance."-/ +@[to_additive "The definition `to_H_space` is not an instance because it comes together with a +multiplicative version which would lead to a diamond since a topological field would inherit two +`H_space` structures, one from the `mul_one_class` and one from the `add_zero_class`. In the case +of an additive group, we make `topological_group.H_space` an instance."] +definition to_H_space (M : Type u) [mul_one_class M] [topological_space M] + [has_continuous_mul M] : H_space M := { Hmul := ⟨function.uncurry has_mul.mul, continuous_mul⟩, e := 1, Hmul_e_e := one_mul 1, e_Hmul := (homotopy_rel.refl _ _).cast rfl (by {ext1, apply one_mul}), Hmul_e := (homotopy_rel.refl _ _).cast rfl (by {ext1, apply mul_one}) } +@[priority 600, to_additive] instance H_space (G : Type u) + [topological_space G] [group G] [topological_group G] : H_space G := to_H_space G + lemma one_eq_H_space_e {G : Type u} [topological_space G] [group G] [topological_group G] : (1 : G) = H_space.e := rfl -/- -In the following example we see that the `H-space` structure on the product of two topological -groups is definitionally equally to the product `H-space`-structure of the two groups. --/ +/- In the following example we see that the `H-space` structure on the product of two topological +groups is definitionally equally to the product `H-space`-structure of the two groups.-/ + example {G G' : Type u} [topological_space G] [group G] [topological_group G] [topological_space G'] [group G'] [topological_group G'] : - topological_group.H_space (G × G') = H_space.prod G G' := rfl - + to_H_space (G × G') = H_space.prod G G' := rfl end topological_group From 837f72de63ad6cd96519cde5f1ffd5ed8d280ad0 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 29 May 2023 17:56:41 +0000 Subject: [PATCH 1110/1291] chore(analysis/normed/group/add_torsor): nnnorm lemmas (#18997) For every `dist`/`norm` lemma in these files, this adds the corresponding `nndist`/`nnnorm` lemma. --- src/analysis/normed/group/add_torsor.lean | 27 ++++++++-- src/analysis/normed_space/add_torsor.lean | 60 +++++++++++++++++++++++ 2 files changed, 83 insertions(+), 4 deletions(-) diff --git a/src/analysis/normed/group/add_torsor.lean b/src/analysis/normed/group/add_torsor.lean index 169615af036ed..6642720beafb0 100644 --- a/src/analysis/normed/group/add_torsor.lean +++ b/src/analysis/normed/group/add_torsor.lean @@ -68,12 +68,18 @@ lemma, it is necessary to have `V` as an explicit argument; otherwise `rw dist_eq_norm_vsub` sometimes doesn't work. -/ lemma dist_eq_norm_vsub (x y : P) : dist x y = ‖x -ᵥ y‖ := normed_add_torsor.dist_eq_norm' x y +lemma nndist_eq_nnnorm_vsub (x y : P) : nndist x y = ‖x -ᵥ y‖₊ := +nnreal.eq $ dist_eq_norm_vsub V x y + /-- The distance equals the norm of subtracting two points. In this lemma, it is necessary to have `V` as an explicit argument; otherwise `rw dist_eq_norm_vsub'` sometimes doesn't work. -/ lemma dist_eq_norm_vsub' (x y : P) : dist x y = ‖y -ᵥ x‖ := (dist_comm _ _).trans (dist_eq_norm_vsub _ _ _) +lemma nndist_eq_nnnorm_vsub' (x y : P) : nndist x y = ‖y -ᵥ x‖₊ := +nnreal.eq $ dist_eq_norm_vsub' V x y + end lemma dist_vadd_cancel_left (v : V) (x y : P) : @@ -84,12 +90,22 @@ dist_vadd _ _ _ dist (v₁ +ᵥ x) (v₂ +ᵥ x) = dist v₁ v₂ := by rw [dist_eq_norm_vsub V, dist_eq_norm, vadd_vsub_vadd_cancel_right] +@[simp] lemma nndist_vadd_cancel_right (v₁ v₂ : V) (x : P) : + nndist (v₁ +ᵥ x) (v₂ +ᵥ x) = nndist v₁ v₂ := +nnreal.eq $ dist_vadd_cancel_right _ _ _ + @[simp] lemma dist_vadd_left (v : V) (x : P) : dist (v +ᵥ x) x = ‖v‖ := by simp [dist_eq_norm_vsub V _ x] +@[simp] lemma nndist_vadd_left (v : V) (x : P) : nndist (v +ᵥ x) x = ‖v‖₊ := +nnreal.eq $ dist_vadd_left _ _ + @[simp] lemma dist_vadd_right (v : V) (x : P) : dist x (v +ᵥ x) = ‖v‖ := by rw [dist_comm, dist_vadd_left] +@[simp] lemma nndist_vadd_right (v : V) (x : P) : nndist x (v +ᵥ x) = ‖v‖₊ := +nnreal.eq $ dist_vadd_right _ _ + /-- Isometry between the tangent space `V` of a (semi)normed add torsor `P` and `P` given by addition/subtraction of `x : P`. -/ @[simps] def isometry_equiv.vadd_const (x : P) : V ≃ᵢ P := @@ -108,19 +124,22 @@ subtraction from `x : P`. -/ @[simp] lemma dist_vsub_cancel_right (x y z : P) : dist (x -ᵥ z) (y -ᵥ z) = dist x y := (isometry_equiv.vadd_const z).symm.dist_eq x y +@[simp] lemma nndist_vsub_cancel_right (x y z : P) : nndist (x -ᵥ z) (y -ᵥ z) = nndist x y := +nnreal.eq $ dist_vsub_cancel_right _ _ _ + lemma dist_vadd_vadd_le (v v' : V) (p p' : P) : dist (v +ᵥ p) (v' +ᵥ p') ≤ dist v v' + dist p p' := by simpa using dist_triangle (v +ᵥ p) (v' +ᵥ p) (v' +ᵥ p') +lemma nndist_vadd_vadd_le (v v' : V) (p p' : P) : + nndist (v +ᵥ p) (v' +ᵥ p') ≤ nndist v v' + nndist p p' := +dist_vadd_vadd_le _ _ _ _ + lemma dist_vsub_vsub_le (p₁ p₂ p₃ p₄ : P) : dist (p₁ -ᵥ p₂) (p₃ -ᵥ p₄) ≤ dist p₁ p₃ + dist p₂ p₄ := by { rw [dist_eq_norm, vsub_sub_vsub_comm, dist_eq_norm_vsub V, dist_eq_norm_vsub V], exact norm_sub_le _ _ } -lemma nndist_vadd_vadd_le (v v' : V) (p p' : P) : - nndist (v +ᵥ p) (v' +ᵥ p') ≤ nndist v v' + nndist p p' := -by simp only [← nnreal.coe_le_coe, nnreal.coe_add, ← dist_nndist, dist_vadd_vadd_le] - lemma nndist_vsub_vsub_le (p₁ p₂ p₃ p₄ : P) : nndist (p₁ -ᵥ p₂) (p₃ -ᵥ p₄) ≤ nndist p₁ p₃ + nndist p₂ p₄ := by simp only [← nnreal.coe_le_coe, nnreal.coe_add, ← dist_nndist, dist_vsub_vsub_le] diff --git a/src/analysis/normed_space/add_torsor.lean b/src/analysis/normed_space/add_torsor.lean index 001c01845a245..21887171303ef 100644 --- a/src/analysis/normed_space/add_torsor.lean +++ b/src/analysis/normed_space/add_torsor.lean @@ -46,10 +46,18 @@ include V dist p₁ (homothety p₁ c p₂) = ‖c‖ * dist p₁ p₂ := by simp [homothety_def, norm_smul, ← dist_eq_norm_vsub, dist_comm] +@[simp] lemma nndist_center_homothety (p₁ p₂ : P) (c : 𝕜) : + nndist p₁ (homothety p₁ c p₂) = ‖c‖₊ * nndist p₁ p₂ := +nnreal.eq $ dist_center_homothety _ _ _ + @[simp] lemma dist_homothety_center (p₁ p₂ : P) (c : 𝕜) : dist (homothety p₁ c p₂) p₁ = ‖c‖ * dist p₁ p₂ := by rw [dist_comm, dist_center_homothety] +@[simp] lemma nndist_homothety_center (p₁ p₂ : P) (c : 𝕜) : + nndist (homothety p₁ c p₂) p₁ = ‖c‖₊ * nndist p₁ p₂ := +nnreal.eq $ dist_homothety_center _ _ _ + @[simp] lemma dist_line_map_line_map (p₁ p₂ : P) (c₁ c₂ : 𝕜) : dist (line_map p₁ p₂ c₁) (line_map p₁ p₂ c₂) = dist c₁ c₂ * dist p₁ p₂ := begin @@ -58,6 +66,10 @@ begin vsub_eq_sub], end +@[simp] lemma nndist_line_map_line_map (p₁ p₂ : P) (c₁ c₂ : 𝕜) : + nndist (line_map p₁ p₂ c₁) (line_map p₁ p₂ c₂) = nndist c₁ c₂ * nndist p₁ p₂ := +nnreal.eq $ dist_line_map_line_map _ _ _ _ + lemma lipschitz_with_line_map (p₁ p₂ : P) : lipschitz_with (nndist p₁ p₂) (line_map p₁ p₂ : 𝕜 → P) := lipschitz_with.of_dist_le_mul $ λ c₁ c₂, @@ -67,26 +79,50 @@ lipschitz_with.of_dist_le_mul $ λ c₁ c₂, dist (line_map p₁ p₂ c) p₁ = ‖c‖ * dist p₁ p₂ := by simpa only [line_map_apply_zero, dist_zero_right] using dist_line_map_line_map p₁ p₂ c 0 +@[simp] lemma nndist_line_map_left (p₁ p₂ : P) (c : 𝕜) : + nndist (line_map p₁ p₂ c) p₁ = ‖c‖₊ * nndist p₁ p₂ := +nnreal.eq $ dist_line_map_left _ _ _ + @[simp] lemma dist_left_line_map (p₁ p₂ : P) (c : 𝕜) : dist p₁ (line_map p₁ p₂ c) = ‖c‖ * dist p₁ p₂ := (dist_comm _ _).trans (dist_line_map_left _ _ _) +@[simp] lemma nndist_left_line_map (p₁ p₂ : P) (c : 𝕜) : + nndist p₁ (line_map p₁ p₂ c) = ‖c‖₊ * nndist p₁ p₂ := +nnreal.eq $ dist_left_line_map _ _ _ + @[simp] lemma dist_line_map_right (p₁ p₂ : P) (c : 𝕜) : dist (line_map p₁ p₂ c) p₂ = ‖1 - c‖ * dist p₁ p₂ := by simpa only [line_map_apply_one, dist_eq_norm'] using dist_line_map_line_map p₁ p₂ c 1 +@[simp] lemma nndist_line_map_right (p₁ p₂ : P) (c : 𝕜) : + nndist (line_map p₁ p₂ c) p₂ = ‖1 - c‖₊ * nndist p₁ p₂ := +nnreal.eq $ dist_line_map_right _ _ _ + @[simp] lemma dist_right_line_map (p₁ p₂ : P) (c : 𝕜) : dist p₂ (line_map p₁ p₂ c) = ‖1 - c‖ * dist p₁ p₂ := (dist_comm _ _).trans (dist_line_map_right _ _ _) +@[simp] lemma nndist_right_line_map (p₁ p₂ : P) (c : 𝕜) : + nndist p₂ (line_map p₁ p₂ c) = ‖1 - c‖₊ * nndist p₁ p₂ := +nnreal.eq $ dist_right_line_map _ _ _ + @[simp] lemma dist_homothety_self (p₁ p₂ : P) (c : 𝕜) : dist (homothety p₁ c p₂) p₂ = ‖1 - c‖ * dist p₁ p₂ := by rw [homothety_eq_line_map, dist_line_map_right] +@[simp] lemma nndist_homothety_self (p₁ p₂ : P) (c : 𝕜) : + nndist (homothety p₁ c p₂) p₂ = ‖1 - c‖₊ * nndist p₁ p₂ := +nnreal.eq $ dist_homothety_self _ _ _ + @[simp] lemma dist_self_homothety (p₁ p₂ : P) (c : 𝕜) : dist p₂ (homothety p₁ c p₂) = ‖1 - c‖ * dist p₁ p₂ := by rw [dist_comm, dist_homothety_self] +@[simp] lemma nndist_self_homothety (p₁ p₂ : P) (c : 𝕜) : + nndist p₂ (homothety p₁ c p₂) = ‖1 - c‖₊ * nndist p₁ p₂ := +nnreal.eq $ dist_self_homothety _ _ _ + section invertible_two variables [invertible (2:𝕜)] @@ -95,18 +131,34 @@ variables [invertible (2:𝕜)] dist p₁ (midpoint 𝕜 p₁ p₂) = ‖(2:𝕜)‖⁻¹ * dist p₁ p₂ := by rw [midpoint, dist_comm, dist_line_map_left, inv_of_eq_inv, ← norm_inv] +@[simp] lemma nndist_left_midpoint (p₁ p₂ : P) : + nndist p₁ (midpoint 𝕜 p₁ p₂) = ‖(2:𝕜)‖₊⁻¹ * nndist p₁ p₂ := +nnreal.eq $ dist_left_midpoint _ _ + @[simp] lemma dist_midpoint_left (p₁ p₂ : P) : dist (midpoint 𝕜 p₁ p₂) p₁ = ‖(2:𝕜)‖⁻¹ * dist p₁ p₂ := by rw [dist_comm, dist_left_midpoint] +@[simp] lemma nndist_midpoint_left (p₁ p₂ : P) : + nndist (midpoint 𝕜 p₁ p₂) p₁ = ‖(2:𝕜)‖₊⁻¹ * nndist p₁ p₂ := +nnreal.eq $ dist_midpoint_left _ _ + @[simp] lemma dist_midpoint_right (p₁ p₂ : P) : dist (midpoint 𝕜 p₁ p₂) p₂ = ‖(2:𝕜)‖⁻¹ * dist p₁ p₂ := by rw [midpoint_comm, dist_midpoint_left, dist_comm] +@[simp] lemma nndist_midpoint_right (p₁ p₂ : P) : + nndist (midpoint 𝕜 p₁ p₂) p₂ = ‖(2:𝕜)‖₊⁻¹ * nndist p₁ p₂ := +nnreal.eq $ dist_midpoint_right _ _ + @[simp] lemma dist_right_midpoint (p₁ p₂ : P) : dist p₂ (midpoint 𝕜 p₁ p₂) = ‖(2:𝕜)‖⁻¹ * dist p₁ p₂ := by rw [dist_comm, dist_midpoint_right] +@[simp] lemma nndist_right_midpoint (p₁ p₂ : P) : + nndist p₂ (midpoint 𝕜 p₁ p₂) = ‖(2:𝕜)‖₊⁻¹ * nndist p₁ p₂ := +nnreal.eq $ dist_right_midpoint _ _ + lemma dist_midpoint_midpoint_le' (p₁ p₂ p₃ p₄ : P) : dist (midpoint 𝕜 p₁ p₂) (midpoint 𝕜 p₃ p₄) ≤ (dist p₁ p₃ + dist p₂ p₄) / ‖(2 : 𝕜)‖ := begin @@ -116,6 +168,10 @@ begin exact div_le_div_of_le_of_nonneg (norm_add_le _ _) (norm_nonneg _), end +lemma nndist_midpoint_midpoint_le' (p₁ p₂ p₃ p₄ : P) : + nndist (midpoint 𝕜 p₁ p₂) (midpoint 𝕜 p₃ p₄) ≤ (nndist p₁ p₃ + nndist p₂ p₄) / ‖(2 : 𝕜)‖₊ := +dist_midpoint_midpoint_le' _ _ _ _ + end invertible_two omit V @@ -160,6 +216,10 @@ lemma dist_midpoint_midpoint_le (p₁ p₂ p₃ p₄ : V) : dist (midpoint ℝ p₁ p₂) (midpoint ℝ p₃ p₄) ≤ (dist p₁ p₃ + dist p₂ p₄) / 2 := by simpa using dist_midpoint_midpoint_le' p₁ p₂ p₃ p₄ +lemma nndist_midpoint_midpoint_le (p₁ p₂ p₃ p₄ : V) : + nndist (midpoint ℝ p₁ p₂) (midpoint ℝ p₃ p₄) ≤ (nndist p₁ p₃ + nndist p₂ p₄) / 2 := +dist_midpoint_midpoint_le _ _ _ _ + include V W /-- A continuous map between two normed affine spaces is an affine map provided that From bd65478311e4dfd41f48bf38c7e3b02fb75d0163 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Tue, 30 May 2023 03:49:43 +0000 Subject: [PATCH 1111/1291] =?UTF-8?q?chore(linear=5Falgebra/alternating):?= =?UTF-8?q?=20make=20`=CE=B9`=20an=20explicit=20arg=20of=20`alternating=5F?= =?UTF-8?q?map.const=5Fof=5Fis=5Fempty`=20(#19123)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit While for general multilinear maps one can deduce it from the type of `E : ι -> Type*`, this doesn't work for alternating maps. --- src/analysis/inner_product_space/orientation.lean | 2 +- src/linear_algebra/alternating.lean | 4 +++- src/linear_algebra/determinant.lean | 2 +- 3 files changed, 5 insertions(+), 3 deletions(-) diff --git a/src/analysis/inner_product_space/orientation.lean b/src/analysis/inner_product_space/orientation.lean index d962ce8bfe730..7066e98cf60d5 100644 --- a/src/analysis/inner_product_space/orientation.lean +++ b/src/analysis/inner_product_space/orientation.lean @@ -181,7 +181,7 @@ alternating form uniquely defined by compatibility with the orientation and inne begin classical, unfreezingI { cases n }, - { let opos : alternating_map ℝ E ℝ (fin 0) := alternating_map.const_of_is_empty ℝ E (1:ℝ), + { let opos : alternating_map ℝ E ℝ (fin 0) := alternating_map.const_of_is_empty ℝ E (fin 0) (1:ℝ), exact o.eq_or_eq_neg_of_is_empty.by_cases (λ _, opos) (λ _, -opos) }, { exact (o.fin_orthonormal_basis n.succ_pos _i.out).to_basis.det } end diff --git a/src/linear_algebra/alternating.lean b/src/linear_algebra/alternating.lean index c0a1adb614edd..0fa1eccdefb48 100644 --- a/src/linear_algebra/alternating.lean +++ b/src/linear_algebra/alternating.lean @@ -337,6 +337,8 @@ def of_subsingleton [subsingleton ι] (i : ι) : alternating_map R M M ι := map_eq_zero_of_eq' := λ v i j hv hij, (hij $ subsingleton.elim _ _).elim, ..multilinear_map.of_subsingleton R M i } +variable (ι) + /-- The constant map is alternating when `ι` is empty. -/ @[simps {fully_applied := ff}] def const_of_is_empty [is_empty ι] (m : N) : alternating_map R M N ι := @@ -1102,7 +1104,7 @@ end to an empty family. -/ @[simps] def const_linear_equiv_of_is_empty [is_empty ι] : N'' ≃ₗ[R'] alternating_map R' M'' N'' ι := -{ to_fun := alternating_map.const_of_is_empty R' M'', +{ to_fun := alternating_map.const_of_is_empty R' M'' ι, map_add' := λ x y, rfl, map_smul' := λ t x, rfl, inv_fun := λ f, f 0, diff --git a/src/linear_algebra/determinant.lean b/src/linear_algebra/determinant.lean index c7a57875c72af..80d805eccb482 100644 --- a/src/linear_algebra/determinant.lean +++ b/src/linear_algebra/determinant.lean @@ -480,7 +480,7 @@ lemma basis.det_apply (v : ι → M) : e.det v = det (e.to_matrix v) := rfl lemma basis.det_self : e.det e = 1 := by simp [e.det_apply] -@[simp] lemma basis.det_is_empty [is_empty ι] : e.det = alternating_map.const_of_is_empty R M 1 := +@[simp] lemma basis.det_is_empty [is_empty ι] : e.det = alternating_map.const_of_is_empty R M ι 1 := begin ext v, exact matrix.det_is_empty, From f60c6087a7275b72d5db3c5a1d0e19e35a429c0a Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Tue, 30 May 2023 08:50:51 +0000 Subject: [PATCH 1112/1291] chore(*): add mathlib4 synchronization comments (#19115) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.category.Group.biproducts` * `algebra.category.Module.monoidal.basic` * `algebra.category.Ring.adjunctions` * `algebra.monoid_algebra.grading` * `analysis.calculus.deriv.add` * `analysis.calculus.deriv.basic` * `analysis.calculus.deriv.comp` * `analysis.calculus.deriv.inverse` * `analysis.calculus.deriv.linear` * `analysis.calculus.deriv.mul` * `analysis.calculus.deriv.polynomial` * `analysis.calculus.deriv.pow` * `analysis.calculus.deriv.prod` * `analysis.calculus.deriv.slope` * `analysis.calculus.deriv.star` * `analysis.calculus.deriv.support` * `analysis.calculus.fderiv_measurable` * `analysis.calculus.local_extr` * `analysis.convex.cone.dual` * `analysis.inner_product_space.lax_milgram` * `analysis.inner_product_space.pi_L2` * `category_theory.bicategory.coherence` * `category_theory.limits.fubini` * `category_theory.monoidal.coherence_lemmas` * `geometry.euclidean.sphere.basic` * `geometry.euclidean.sphere.power` * `geometry.euclidean.sphere.second_inter` * `group_theory.complement` * `group_theory.sylow` * `linear_algebra.matrix.hermitian` * `measure_theory.constructions.prod.basic` * `measure_theory.function.lp_order` * `measure_theory.function.lp_seminorm` * `measure_theory.function.lp_space` * `measure_theory.measure.complex` * `measure_theory.measure.vector_measure` * `number_theory.padics.ring_homs` * `ring_theory.ideal.cotangent` * `topology.metric_space.metrizable_uniformity` * `topology.sheaves.forget` * `topology.sheaves.sheaf_condition.unique_gluing` * `topology.sheaves.sheaf_of_functions` --- src/algebra/category/Group/biproducts.lean | 3 +++ src/algebra/category/Module/monoidal/basic.lean | 3 +++ src/algebra/category/Ring/adjunctions.lean | 3 +++ src/algebra/monoid_algebra/grading.lean | 3 +++ src/analysis/calculus/deriv/add.lean | 3 +++ src/analysis/calculus/deriv/basic.lean | 3 +++ src/analysis/calculus/deriv/comp.lean | 3 +++ src/analysis/calculus/deriv/inverse.lean | 3 +++ src/analysis/calculus/deriv/linear.lean | 3 +++ src/analysis/calculus/deriv/mul.lean | 3 +++ src/analysis/calculus/deriv/polynomial.lean | 3 +++ src/analysis/calculus/deriv/pow.lean | 3 +++ src/analysis/calculus/deriv/prod.lean | 3 +++ src/analysis/calculus/deriv/slope.lean | 3 +++ src/analysis/calculus/deriv/star.lean | 3 +++ src/analysis/calculus/deriv/support.lean | 3 +++ src/analysis/calculus/fderiv_measurable.lean | 3 +++ src/analysis/calculus/local_extr.lean | 3 +++ src/analysis/convex/cone/dual.lean | 3 +++ src/analysis/inner_product_space/lax_milgram.lean | 3 +++ src/analysis/inner_product_space/pi_L2.lean | 3 +++ src/category_theory/bicategory/coherence.lean | 3 +++ src/category_theory/limits/fubini.lean | 3 +++ src/category_theory/monoidal/coherence_lemmas.lean | 3 +++ src/geometry/euclidean/sphere/basic.lean | 3 +++ src/geometry/euclidean/sphere/power.lean | 3 +++ src/geometry/euclidean/sphere/second_inter.lean | 3 +++ src/group_theory/complement.lean | 3 +++ src/group_theory/sylow.lean | 3 +++ src/linear_algebra/matrix/hermitian.lean | 3 +++ src/measure_theory/constructions/prod/basic.lean | 3 +++ src/measure_theory/function/lp_order.lean | 3 +++ src/measure_theory/function/lp_seminorm.lean | 3 +++ src/measure_theory/function/lp_space.lean | 3 +++ src/measure_theory/measure/complex.lean | 3 +++ src/measure_theory/measure/vector_measure.lean | 3 +++ src/number_theory/padics/ring_homs.lean | 3 +++ src/ring_theory/ideal/cotangent.lean | 3 +++ src/topology/metric_space/metrizable_uniformity.lean | 3 +++ src/topology/sheaves/forget.lean | 3 +++ src/topology/sheaves/sheaf_condition/unique_gluing.lean | 3 +++ src/topology/sheaves/sheaf_of_functions.lean | 3 +++ 42 files changed, 126 insertions(+) diff --git a/src/algebra/category/Group/biproducts.lean b/src/algebra/category/Group/biproducts.lean index 7616b1378093f..c11671aab003c 100644 --- a/src/algebra/category/Group/biproducts.lean +++ b/src/algebra/category/Group/biproducts.lean @@ -10,6 +10,9 @@ import algebra.category.Group.limits /-! # The category of abelian groups has finite biproducts + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ open category_theory diff --git a/src/algebra/category/Module/monoidal/basic.lean b/src/algebra/category/Module/monoidal/basic.lean index 670fb7c145ad0..83862f8481145 100644 --- a/src/algebra/category/Module/monoidal/basic.lean +++ b/src/algebra/category/Module/monoidal/basic.lean @@ -11,6 +11,9 @@ import category_theory.monoidal.linear /-! # The monoidal category structure on R-modules +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Mostly this uses existing machinery in `linear_algebra.tensor_product`. We just need to provide a few small missing pieces to build the `monoidal_category` instance. diff --git a/src/algebra/category/Ring/adjunctions.lean b/src/algebra/category/Ring/adjunctions.lean index d1d6809d85fea..166a5e3c14e97 100644 --- a/src/algebra/category/Ring/adjunctions.lean +++ b/src/algebra/category/Ring/adjunctions.lean @@ -7,6 +7,9 @@ import algebra.category.Ring.basic import data.mv_polynomial.comm_ring /-! +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Multivariable polynomials on a type is the left adjoint of the forgetful functor from commutative rings to types. -/ diff --git a/src/algebra/monoid_algebra/grading.lean b/src/algebra/monoid_algebra/grading.lean index 3fa0098a8a862..d0acc07e315b1 100644 --- a/src/algebra/monoid_algebra/grading.lean +++ b/src/algebra/monoid_algebra/grading.lean @@ -11,6 +11,9 @@ import ring_theory.graded_algebra.basic /-! # Internal grading of an `add_monoid_algebra` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we show that an `add_monoid_algebra` has an internal direct sum structure. ## Main results diff --git a/src/analysis/calculus/deriv/add.lean b/src/analysis/calculus/deriv/add.lean index 70f4e35833266..c199f455536e6 100644 --- a/src/analysis/calculus/deriv/add.lean +++ b/src/analysis/calculus/deriv/add.lean @@ -9,6 +9,9 @@ import analysis.calculus.fderiv.add /-! # One-dimensional derivatives of sums etc +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove formulas about derivatives of `f + g`, `-f`, `f - g`, and `∑ i, f i x` for functions from the base field to a normed space over this field. diff --git a/src/analysis/calculus/deriv/basic.lean b/src/analysis/calculus/deriv/basic.lean index cf6c1d7c00b02..f80d23c7f919e 100644 --- a/src/analysis/calculus/deriv/basic.lean +++ b/src/analysis/calculus/deriv/basic.lean @@ -9,6 +9,9 @@ import analysis.calculus.fderiv.basic # One-dimensional derivatives +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the derivative of a function `f : 𝕜 → F` where `𝕜` is a normed field and `F` is a normed space over this field. The derivative of such a function `f` at a point `x` is given by an element `f' : F`. diff --git a/src/analysis/calculus/deriv/comp.lean b/src/analysis/calculus/deriv/comp.lean index 1fdb3fd794b61..44c390d99966f 100644 --- a/src/analysis/calculus/deriv/comp.lean +++ b/src/analysis/calculus/deriv/comp.lean @@ -10,6 +10,9 @@ import analysis.calculus.fderiv.restrict_scalars /-! # One-dimensional derivatives of compositions of functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove the chain rule for the following cases: * `has_deriv_at.comp` etc: `f : 𝕜' → 𝕜'` composed with `g : 𝕜 → 𝕜'`; diff --git a/src/analysis/calculus/deriv/inverse.lean b/src/analysis/calculus/deriv/inverse.lean index 4e21ac6ced7f9..6d1e9842d9a79 100644 --- a/src/analysis/calculus/deriv/inverse.lean +++ b/src/analysis/calculus/deriv/inverse.lean @@ -9,6 +9,9 @@ import analysis.calculus.fderiv.equiv /-! # Inverse function theorem - the easy half +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove that `g' (f x) = (f' x)⁻¹` provided that `f` is strictly differentiable at `x`, `f' x ≠ 0`, and `g` is a local left inverse of `f` that is continuous at `f x`. This is the easy half of the inverse function theorem: the harder half states that `g` exists. diff --git a/src/analysis/calculus/deriv/linear.lean b/src/analysis/calculus/deriv/linear.lean index 4b6ab842f9c5d..45198e21eb2b0 100644 --- a/src/analysis/calculus/deriv/linear.lean +++ b/src/analysis/calculus/deriv/linear.lean @@ -9,6 +9,9 @@ import analysis.calculus.fderiv.linear /-! # Derivatives of continuous linear maps from the base field +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove that `f : 𝕜 →L[𝕜] E` (or `f : 𝕜 →ₗ[𝕜] E`) has derivative `f 1`. For a more detailed overview of one-dimensional derivatives in mathlib, see the module docstring of diff --git a/src/analysis/calculus/deriv/mul.lean b/src/analysis/calculus/deriv/mul.lean index c99cc0e0b637b..3a21aef1819eb 100644 --- a/src/analysis/calculus/deriv/mul.lean +++ b/src/analysis/calculus/deriv/mul.lean @@ -10,6 +10,9 @@ import analysis.calculus.fderiv.add /-! # Derivative of `f x * g x` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove formulas for `(f x * g x)'` and `(f x • g x)'`. For a more detailed overview of one-dimensional derivatives in mathlib, see the module docstring of diff --git a/src/analysis/calculus/deriv/polynomial.lean b/src/analysis/calculus/deriv/polynomial.lean index 089e604c39f31..9f1d651ae5b2c 100644 --- a/src/analysis/calculus/deriv/polynomial.lean +++ b/src/analysis/calculus/deriv/polynomial.lean @@ -11,6 +11,9 @@ import data.polynomial.derivative /-! # Derivatives of polynomials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove that derivatives of polynomials in the analysis sense agree with their derivatives in the algebraic sense. diff --git a/src/analysis/calculus/deriv/pow.lean b/src/analysis/calculus/deriv/pow.lean index d1c0b5796657c..8a35ec8a49940 100644 --- a/src/analysis/calculus/deriv/pow.lean +++ b/src/analysis/calculus/deriv/pow.lean @@ -9,6 +9,9 @@ import analysis.calculus.deriv.comp /-! # Derivative of `(f x) ^ n`, `n : ℕ` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove that `(x ^ n)' = n * x ^ (n - 1)`, where `n` is a natural number. For a more detailed overview of one-dimensional derivatives in mathlib, see the module docstring of diff --git a/src/analysis/calculus/deriv/prod.lean b/src/analysis/calculus/deriv/prod.lean index 2ab31ee57d62c..70b478ea5c399 100644 --- a/src/analysis/calculus/deriv/prod.lean +++ b/src/analysis/calculus/deriv/prod.lean @@ -9,6 +9,9 @@ import analysis.calculus.fderiv.prod /-! # Derivatives of functions taking values in product types +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove lemmas about derivatives of functions `f : 𝕜 → E × F` and of functions `f : 𝕜 → (Π i, E i)`. diff --git a/src/analysis/calculus/deriv/slope.lean b/src/analysis/calculus/deriv/slope.lean index 7f883efca6727..659b3e4190984 100644 --- a/src/analysis/calculus/deriv/slope.lean +++ b/src/analysis/calculus/deriv/slope.lean @@ -9,6 +9,9 @@ import linear_algebra.affine_space.slope /-! # Derivative as the limit of the slope +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we relate the derivative of a function with its definition from a standard undergraduate course as the limit of the slope `(f y - f x) / (y - x)` as `y` tends to `𝓝[≠] x`. Since we are talking about functions taking values in a normed space instead of the base field, we diff --git a/src/analysis/calculus/deriv/star.lean b/src/analysis/calculus/deriv/star.lean index 6f9ad983dd1fd..1f75c711e9ad2 100644 --- a/src/analysis/calculus/deriv/star.lean +++ b/src/analysis/calculus/deriv/star.lean @@ -9,6 +9,9 @@ import analysis.calculus.fderiv.star /-! # Star operations on derivatives +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains the usual formulas (and existence assertions) for the derivative of the star operation. Note that these only apply when the field that the derivative is respect to has a trivial star operation; which as should be expected rules out `𝕜 = ℂ`. diff --git a/src/analysis/calculus/deriv/support.lean b/src/analysis/calculus/deriv/support.lean index cbf46e0ebc3fb..adb9f0637a7ad 100644 --- a/src/analysis/calculus/deriv/support.lean +++ b/src/analysis/calculus/deriv/support.lean @@ -8,6 +8,9 @@ import analysis.calculus.deriv.basic /-! # Support of the derivative of a function +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove that the (topological) support of a function includes the support of its derivative. As a corollary, we show that the derivative of a function with compact support has compact support. diff --git a/src/analysis/calculus/fderiv_measurable.lean b/src/analysis/calculus/fderiv_measurable.lean index 562df4994997c..d3471ac69dd07 100644 --- a/src/analysis/calculus/fderiv_measurable.lean +++ b/src/analysis/calculus/fderiv_measurable.lean @@ -10,6 +10,9 @@ import measure_theory.function.strongly_measurable.basic /-! # Derivative is measurable +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove that the derivative of any function with complete codomain is a measurable function. Namely, we prove: diff --git a/src/analysis/calculus/local_extr.lean b/src/analysis/calculus/local_extr.lean index f0fb3bcf9859d..4cdcfeacb43e7 100644 --- a/src/analysis/calculus/local_extr.lean +++ b/src/analysis/calculus/local_extr.lean @@ -10,6 +10,9 @@ import topology.algebra.polynomial /-! # Local extrema of smooth functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions In a real normed space `E` we define `pos_tangent_cone_at (s : set E) (x : E)`. diff --git a/src/analysis/convex/cone/dual.lean b/src/analysis/convex/cone/dual.lean index a80d36f8cff61..c8d05c0bce19a 100644 --- a/src/analysis/convex/cone/dual.lean +++ b/src/analysis/convex/cone/dual.lean @@ -9,6 +9,9 @@ import analysis.inner_product_space.projection /-! # Convex cones in inner product spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define `set.inner_dual_cone` to be the cone consisting of all points `y` such that for all points `x` in a given set `0 ≤ ⟪ x, y ⟫`. diff --git a/src/analysis/inner_product_space/lax_milgram.lean b/src/analysis/inner_product_space/lax_milgram.lean index c6f4f84191f2a..27354056f5634 100644 --- a/src/analysis/inner_product_space/lax_milgram.lean +++ b/src/analysis/inner_product_space/lax_milgram.lean @@ -12,6 +12,9 @@ import topology.metric_space.antilipschitz /-! # The Lax-Milgram Theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We consider an Hilbert space `V` over `ℝ` equipped with a bounded bilinear form `B : V →L[ℝ] V →L[ℝ] ℝ`. diff --git a/src/analysis/inner_product_space/pi_L2.lean b/src/analysis/inner_product_space/pi_L2.lean index 93730db0aff12..ccad358197ccc 100644 --- a/src/analysis/inner_product_space/pi_L2.lean +++ b/src/analysis/inner_product_space/pi_L2.lean @@ -11,6 +11,9 @@ import linear_algebra.unitary_group /-! # `L²` inner product space structure on finite products of inner product spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The `L²` norm on a finite product of inner product spaces is compatible with an inner product $$ \langle x, y\rangle = \sum \langle x_i, y_i \rangle. diff --git a/src/category_theory/bicategory/coherence.lean b/src/category_theory/bicategory/coherence.lean index 0b3faa06ef847..9cf7ea46cddf7 100644 --- a/src/category_theory/bicategory/coherence.lean +++ b/src/category_theory/bicategory/coherence.lean @@ -10,6 +10,9 @@ import category_theory.bicategory.locally_discrete /-! # The coherence theorem for bicategories +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we prove the coherence theorem for bicategories, stated in the following form: the free bicategory over any quiver is locally thin. diff --git a/src/category_theory/limits/fubini.lean b/src/category_theory/limits/fubini.lean index 3dad12f3e9385..5c6ca492bcde5 100644 --- a/src/category_theory/limits/fubini.lean +++ b/src/category_theory/limits/fubini.lean @@ -10,6 +10,9 @@ import category_theory.functor.currying /-! # A Fubini theorem for categorical limits +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We prove that $lim_{J × K} G = lim_J (lim_K G(j, -))$ for a functor `G : J × K ⥤ C`, when all the appropriate limits exist. diff --git a/src/category_theory/monoidal/coherence_lemmas.lean b/src/category_theory/monoidal/coherence_lemmas.lean index 8d8ff40a19d3c..0408bb098f37f 100644 --- a/src/category_theory/monoidal/coherence_lemmas.lean +++ b/src/category_theory/monoidal/coherence_lemmas.lean @@ -8,6 +8,9 @@ import category_theory.monoidal.coherence /-! # Lemmas which are consequences of monoidal coherence +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + These lemmas are all proved `by coherence`. ## Future work diff --git a/src/geometry/euclidean/sphere/basic.lean b/src/geometry/euclidean/sphere/basic.lean index a017fe8d65a47..8004575cf2780 100644 --- a/src/geometry/euclidean/sphere/basic.lean +++ b/src/geometry/euclidean/sphere/basic.lean @@ -9,6 +9,9 @@ import geometry.euclidean.basic /-! # Spheres +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines and proves basic results about spheres and cospherical sets of points in Euclidean affine spaces. diff --git a/src/geometry/euclidean/sphere/power.lean b/src/geometry/euclidean/sphere/power.lean index ad14d1991684f..7807c30705a37 100644 --- a/src/geometry/euclidean/sphere/power.lean +++ b/src/geometry/euclidean/sphere/power.lean @@ -9,6 +9,9 @@ import geometry.euclidean.sphere.basic /-! # Power of a point (intersecting chords and secants) +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves basic geometrical results about power of a point (intersecting chords and secants) in spheres in real inner product spaces and Euclidean affine spaces. diff --git a/src/geometry/euclidean/sphere/second_inter.lean b/src/geometry/euclidean/sphere/second_inter.lean index ec52336826a8c..f036d3c70ebdd 100644 --- a/src/geometry/euclidean/sphere/second_inter.lean +++ b/src/geometry/euclidean/sphere/second_inter.lean @@ -8,6 +8,9 @@ import geometry.euclidean.sphere.basic /-! # Second intersection of a sphere and a line +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines and proves basic results about the second intersection of a sphere with a line through a point on that sphere. diff --git a/src/group_theory/complement.lean b/src/group_theory/complement.lean index bb98eb0f21623..fbea4afce7115 100644 --- a/src/group_theory/complement.lean +++ b/src/group_theory/complement.lean @@ -9,6 +9,9 @@ import data.zmod.quotient /-! # Complements +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define the complement of a subgroup. ## Main definitions diff --git a/src/group_theory/sylow.lean b/src/group_theory/sylow.lean index 861713c4025c6..91d962c8014b9 100644 --- a/src/group_theory/sylow.lean +++ b/src/group_theory/sylow.lean @@ -14,6 +14,9 @@ import order.atoms.finite /-! # Sylow theorems +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The Sylow theorems are the following results for every finite group `G` and every prime number `p`. * There exists a Sylow `p`-subgroup of `G`. diff --git a/src/linear_algebra/matrix/hermitian.lean b/src/linear_algebra/matrix/hermitian.lean index b4a8e6b46e124..eb070688eae94 100644 --- a/src/linear_algebra/matrix/hermitian.lean +++ b/src/linear_algebra/matrix/hermitian.lean @@ -7,6 +7,9 @@ import analysis.inner_product_space.pi_L2 /-! # Hermitian matrices +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines hermitian matrices and some basic results about them. See also `is_self_adjoint`, which generalizes this definition to other star rings. diff --git a/src/measure_theory/constructions/prod/basic.lean b/src/measure_theory/constructions/prod/basic.lean index b4ef7bf51ed5b..23adab12a0893 100644 --- a/src/measure_theory/constructions/prod/basic.lean +++ b/src/measure_theory/constructions/prod/basic.lean @@ -11,6 +11,9 @@ import measure_theory.measure.open_pos /-! # The product measure +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define and prove properties about the binary product measure. If `α` and `β` have σ-finite measures `μ` resp. `ν` then `α × β` can be equipped with a σ-finite measure `μ.prod ν` that satisfies `(μ.prod ν) s = ∫⁻ x, ν {y | (x, y) ∈ s} ∂μ`. diff --git a/src/measure_theory/function/lp_order.lean b/src/measure_theory/function/lp_order.lean index fa528399bfcf1..f0e0059b88fbe 100644 --- a/src/measure_theory/function/lp_order.lean +++ b/src/measure_theory/function/lp_order.lean @@ -9,6 +9,9 @@ import measure_theory.function.lp_space /-! # Order related properties of Lp spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ### Results - `Lp E p μ` is an `ordered_add_comm_group` when `E` is a `normed_lattice_add_comm_group`. diff --git a/src/measure_theory/function/lp_seminorm.lean b/src/measure_theory/function/lp_seminorm.lean index 54e5e97930327..7482c420fc669 100644 --- a/src/measure_theory/function/lp_seminorm.lean +++ b/src/measure_theory/function/lp_seminorm.lean @@ -12,6 +12,9 @@ import measure_theory.integral.mean_inequalities /-! # ℒp space +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file describes properties of almost everywhere strongly measurable functions with finite `p`-seminorm, denoted by `snorm f p μ` and defined for `p:ℝ≥0∞` as `0` if `p=0`, `(∫ ‖f a‖^p ∂μ) ^ (1/p)` for `0 < p < ∞` and `ess_sup ‖f‖ μ` for `p=∞`. diff --git a/src/measure_theory/function/lp_space.lean b/src/measure_theory/function/lp_space.lean index 4e52fbbdf66e5..8192c83bd4425 100644 --- a/src/measure_theory/function/lp_space.lean +++ b/src/measure_theory/function/lp_space.lean @@ -10,6 +10,9 @@ import topology.continuous_function.compact /-! # Lp space +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file provides the space `Lp E p μ` as the subtype of elements of `α →ₘ[μ] E` (see ae_eq_fun) such that `snorm f p μ` is finite. For `1 ≤ p`, `snorm` defines a norm and `Lp` is a complete metric space. diff --git a/src/measure_theory/measure/complex.lean b/src/measure_theory/measure/complex.lean index e401513ec3f08..3137af7de2b54 100644 --- a/src/measure_theory/measure/complex.lean +++ b/src/measure_theory/measure/complex.lean @@ -9,6 +9,9 @@ import measure_theory.measure.vector_measure /-! # Complex measure +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves some elementary results about complex measures. In particular, we prove that a complex measure is always in the form `s + it` where `s` and `t` are signed measures. diff --git a/src/measure_theory/measure/vector_measure.lean b/src/measure_theory/measure/vector_measure.lean index f29742363b51f..2bb06f31d6fcb 100644 --- a/src/measure_theory/measure/vector_measure.lean +++ b/src/measure_theory/measure/vector_measure.lean @@ -10,6 +10,9 @@ import analysis.complex.basic # Vector valued measures +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines vector valued measures, which are σ-additive functions from a set to a add monoid `M` such that it maps the empty set and non-measurable sets to zero. In the case that `M = ℝ`, we called the vector measure a signed measure and write `signed_measure α`. diff --git a/src/number_theory/padics/ring_homs.lean b/src/number_theory/padics/ring_homs.lean index 9961c7552ad74..1fb617a8d2683 100644 --- a/src/number_theory/padics/ring_homs.lean +++ b/src/number_theory/padics/ring_homs.lean @@ -11,6 +11,9 @@ import number_theory.padics.padic_integers # Relating `ℤ_[p]` to `zmod (p ^ n)` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we establish connections between the `p`-adic integers $\mathbb{Z}_p$ and the integers modulo powers of `p`, $\mathbb{Z}/p^n\mathbb{Z}$. diff --git a/src/ring_theory/ideal/cotangent.lean b/src/ring_theory/ideal/cotangent.lean index 36b1376b88b13..7743e96ac51d3 100644 --- a/src/ring_theory/ideal/cotangent.lean +++ b/src/ring_theory/ideal/cotangent.lean @@ -12,6 +12,9 @@ import ring_theory.ideal.local_ring /-! # The module `I ⧸ I ^ 2` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we provide special API support for the module `I ⧸ I ^ 2`. The official definition is a quotient module of `I`, but the alternative definition as an ideal of `R ⧸ I ^ 2` is also given, and the two are `R`-equivalent as in `ideal.cotangent_equiv_ideal`. diff --git a/src/topology/metric_space/metrizable_uniformity.lean b/src/topology/metric_space/metrizable_uniformity.lean index 1de408a6d60dc..594f03f511223 100644 --- a/src/topology/metric_space/metrizable_uniformity.lean +++ b/src/topology/metric_space/metrizable_uniformity.lean @@ -8,6 +8,9 @@ import topology.metric_space.metrizable /-! # Metrizable uniform spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove that a uniform space with countably generated uniformity filter is pseudometrizable: there exists a `pseudo_metric_space` structure that generates the same uniformity. The proof follows [Sergey Melikhov, Metrizable uniform spaces][melikhov2011]. diff --git a/src/topology/sheaves/forget.lean b/src/topology/sheaves/forget.lean index fed5cc40b5b12..3fb52ea35b299 100644 --- a/src/topology/sheaves/forget.lean +++ b/src/topology/sheaves/forget.lean @@ -9,6 +9,9 @@ import topology.sheaves.sheaf_condition.equalizer_products /-! # Checking the sheaf condition on the underlying presheaf of types. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + If `G : C ⥤ D` is a functor which reflects isomorphisms and preserves limits (we assume all limits exist in both `C` and `D`), then checking the sheaf condition for a presheaf `F : presheaf C X` diff --git a/src/topology/sheaves/sheaf_condition/unique_gluing.lean b/src/topology/sheaves/sheaf_condition/unique_gluing.lean index e9647e2a86406..2cbe750fc2433 100644 --- a/src/topology/sheaves/sheaf_condition/unique_gluing.lean +++ b/src/topology/sheaves/sheaf_condition/unique_gluing.lean @@ -11,6 +11,9 @@ import category_theory.types /-! # The sheaf condition in terms of unique gluings +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We provide an alternative formulation of the sheaf condition in terms of unique gluings. We work with sheaves valued in a concrete category `C` admitting all limits, whose forgetful diff --git a/src/topology/sheaves/sheaf_of_functions.lean b/src/topology/sheaves/sheaf_of_functions.lean index c4608b98ca192..8b3e8bcc0e660 100644 --- a/src/topology/sheaves/sheaf_of_functions.lean +++ b/src/topology/sheaves/sheaf_of_functions.lean @@ -9,6 +9,9 @@ import topology.sheaves.sheaf_condition.unique_gluing /-! # Sheaf conditions for presheaves of (continuous) functions. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We show that * `Top.presheaf.to_Type_is_sheaf`: not-necessarily-continuous functions into a type form a sheaf * `Top.presheaf.to_Types_is_sheaf`: in fact, these may be dependent functions into a type family From 9a1ffe493625d694e4d38e21637eeba15abd0a0d Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Tue, 30 May 2023 10:24:32 +0000 Subject: [PATCH 1113/1291] refactor(analysis/calculus/darboux): review API (#19125) Add more versions of the main theorem, slightly generalize existing versions. --- src/analysis/calculus/darboux.lean | 116 ++++++++++++++++++++--------- 1 file changed, 81 insertions(+), 35 deletions(-) diff --git a/src/analysis/calculus/darboux.lean b/src/analysis/calculus/darboux.lean index 1d25068eee2f2..529cfb0bf3e91 100644 --- a/src/analysis/calculus/darboux.lean +++ b/src/analysis/calculus/darboux.lean @@ -19,16 +19,14 @@ open_locale topology classical variables {a b : ℝ} {f f' : ℝ → ℝ} /-- **Darboux's theorem**: if `a ≤ b` and `f' a < m < f' b`, then `f' c = m` for some -`c ∈ [a, b]`. -/ +`c ∈ (a, b)`. -/ theorem exists_has_deriv_within_at_eq_of_gt_of_lt (hab : a ≤ b) (hf : ∀ x ∈ (Icc a b), has_deriv_within_at f (f' x) (Icc a b) x) {m : ℝ} (hma : f' a < m) (hmb : m < f' b) : - m ∈ f' '' (Icc a b) := + m ∈ f' '' Ioo a b := begin - have hab' : a < b, - { refine lt_of_le_of_ne hab (λ hab', _), - subst b, - exact lt_asymm hma hmb }, + rcases hab.eq_or_lt with rfl | hab', + { exact (lt_asymm hma hmb).elim }, set g : ℝ → ℝ := λ x, f x - m * x, have hg : ∀ x ∈ Icc a b, has_deriv_within_at g (f' x - m) (Icc a b) x, { intros x hx, @@ -37,73 +35,121 @@ begin from is_compact_Icc.exists_forall_le (nonempty_Icc.2 $ hab) (λ x hx, (hg x hx).continuous_within_at), have cmem' : c ∈ Ioo a b, - { cases eq_or_lt_of_le cmem.1 with hac hac, + { rcases cmem.1.eq_or_lt with rfl | hac, -- Show that `c` can't be equal to `a` - { subst c, - refine absurd (sub_nonneg.1 $ nonneg_of_mul_nonneg_right _ (sub_pos.2 hab')) + { refine absurd (sub_nonneg.1 $ nonneg_of_mul_nonneg_right _ (sub_pos.2 hab')) (not_le_of_lt hma), have : b - a ∈ pos_tangent_cone_at (Icc a b) a, from mem_pos_tangent_cone_at_of_segment_subset (segment_eq_Icc hab ▸ subset.refl _), simpa [-sub_nonneg, -continuous_linear_map.map_sub] using hc.localize.has_fderiv_within_at_nonneg (hg a (left_mem_Icc.2 hab)) this }, - cases eq_or_lt_of_le cmem.2 with hbc hbc, + rcases cmem.2.eq_or_gt with rfl | hcb, -- Show that `c` can't be equal to `b` - { subst c, - refine absurd (sub_nonpos.1 $ nonpos_of_mul_nonneg_right _ (sub_lt_zero.2 hab')) + { refine absurd (sub_nonpos.1 $ nonpos_of_mul_nonneg_right _ (sub_lt_zero.2 hab')) (not_le_of_lt hmb), have : a - b ∈ pos_tangent_cone_at (Icc a b) b, from mem_pos_tangent_cone_at_of_segment_subset (by rw [segment_symm, segment_eq_Icc hab]), simpa [-sub_nonneg, -continuous_linear_map.map_sub] using hc.localize.has_fderiv_within_at_nonneg (hg b (right_mem_Icc.2 hab)) this }, - exact ⟨hac, hbc⟩ }, - use [c, cmem], + exact ⟨hac, hcb⟩ }, + use [c, cmem'], rw [← sub_eq_zero], have : Icc a b ∈ 𝓝 c, by rwa [← mem_interior_iff_mem_nhds, interior_Icc], exact (hc.is_local_min this).has_deriv_at_eq_zero ((hg c cmem).has_deriv_at this) end -/-- **Darboux's theorem**: if `a ≤ b` and `f' a > m > f' b`, then `f' c = m` for some `c ∈ [a, b]`. +/-- **Darboux's theorem**: if `a ≤ b` and `f' b < m < f' a`, then `f' c = m` for some `c ∈ (a, b)`. -/ theorem exists_has_deriv_within_at_eq_of_lt_of_gt (hab : a ≤ b) (hf : ∀ x ∈ (Icc a b), has_deriv_within_at f (f' x) (Icc a b) x) {m : ℝ} (hma : m < f' a) (hmb : f' b < m) : - m ∈ f' '' (Icc a b) := + m ∈ f' '' Ioo a b := let ⟨c, cmem, hc⟩ := exists_has_deriv_within_at_eq_of_gt_of_lt hab (λ x hx, (hf x hx).neg) (neg_lt_neg hma) (neg_lt_neg hmb) in ⟨c, cmem, neg_injective hc⟩ -/-- **Darboux's theorem**: the image of a convex set under `f'` is a convex set. -/ -theorem convex_image_has_deriv_at {s : set ℝ} (hs : convex ℝ s) - (hf : ∀ x ∈ s, has_deriv_at f (f' x) x) : - convex ℝ (f' '' s) := +/-- **Darboux's theorem**: the image of an `ord_connected` set under `f'` is an `ord_connected` +set, `has_deriv_within_at` version. -/ +theorem set.ord_connected.image_has_deriv_within_at {s : set ℝ} (hs : ord_connected s) + (hf : ∀ x ∈ s, has_deriv_within_at f (f' x) s x) : + ord_connected (f' '' s) := begin - refine ord_connected.convex ⟨_⟩, - rintros _ ⟨a, ha, rfl⟩ _ ⟨b, hb, rfl⟩ m ⟨hma, hmb⟩, - cases eq_or_lt_of_le hma with hma hma, - by exact hma ▸ mem_image_of_mem f' ha, - cases eq_or_lt_of_le hmb with hmb hmb, - by exact hmb.symm ▸ mem_image_of_mem f' hb, + apply ord_connected_of_Ioo, + rintros _ ⟨a, ha, rfl⟩ _ ⟨b, hb, rfl⟩ - m ⟨hma, hmb⟩, cases le_total a b with hab hab, - { have : Icc a b ⊆ s, from hs.ord_connected.out ha hb, + { have : Icc a b ⊆ s, from hs.out ha hb, rcases exists_has_deriv_within_at_eq_of_gt_of_lt hab - (λ x hx, (hf x $ this hx).has_deriv_within_at) hma hmb + (λ x hx, (hf x $ this hx).mono this) hma hmb with ⟨c, cmem, hc⟩, - exact ⟨c, this cmem, hc⟩ }, - { have : Icc b a ⊆ s, from hs.ord_connected.out hb ha, + exact ⟨c, this $ Ioo_subset_Icc_self cmem, hc⟩ }, + { have : Icc b a ⊆ s, from hs.out hb ha, rcases exists_has_deriv_within_at_eq_of_lt_of_gt hab - (λ x hx, (hf x $ this hx).has_deriv_within_at) hmb hma + (λ x hx, (hf x $ this hx).mono this) hmb hma with ⟨c, cmem, hc⟩, - exact ⟨c, this cmem, hc⟩ } + exact ⟨c, this $ Ioo_subset_Icc_self cmem, hc⟩ } end +/-- **Darboux's theorem**: the image of an `ord_connected` set under `f'` is an `ord_connected` +set, `deriv_within` version. -/ +theorem set.ord_connected.image_deriv_within {s : set ℝ} (hs : ord_connected s) + (hf : differentiable_on ℝ f s) : + ord_connected (deriv_within f s '' s) := +hs.image_has_deriv_within_at $ λ x hx, (hf x hx).has_deriv_within_at + +/-- **Darboux's theorem**: the image of an `ord_connected` set under `f'` is an `ord_connected` +set, `deriv` version. -/ +theorem set.ord_connected.image_deriv {s : set ℝ} (hs : ord_connected s) + (hf : ∀ x ∈ s, differentiable_at ℝ f x) : + ord_connected (deriv f '' s) := +hs.image_has_deriv_within_at $ λ x hx, (hf x hx).has_deriv_at.has_deriv_within_at + +/-- **Darboux's theorem**: the image of a convex set under `f'` is a convex set, +`has_deriv_within_at` version. -/ +theorem convex.image_has_deriv_within_at {s : set ℝ} (hs : convex ℝ s) + (hf : ∀ x ∈ s, has_deriv_within_at f (f' x) s x) : + convex ℝ (f' '' s) := +(hs.ord_connected.image_has_deriv_within_at hf).convex + +/-- **Darboux's theorem**: the image of a convex set under `f'` is a convex set, +`deriv_within` version. -/ +theorem convex.image_deriv_within {s : set ℝ} (hs : convex ℝ s) + (hf : differentiable_on ℝ f s) : + convex ℝ (deriv_within f s '' s) := +(hs.ord_connected.image_deriv_within hf).convex + +/-- **Darboux's theorem**: the image of a convex set under `f'` is a convex set, +`deriv` version. -/ +theorem convex.image_deriv {s : set ℝ} (hs : convex ℝ s) + (hf : ∀ x ∈ s, differentiable_at ℝ f x) : + convex ℝ (deriv f '' s) := +(hs.ord_connected.image_deriv hf).convex + +/-- **Darboux's theorem**: if `a ≤ b` and `f' a ≤ m ≤ f' b`, then `f' c = m` for some +`c ∈ [a, b]`. -/ +theorem exists_has_deriv_within_at_eq_of_ge_of_le + (hab : a ≤ b) (hf : ∀ x ∈ (Icc a b), has_deriv_within_at f (f' x) (Icc a b) x) + {m : ℝ} (hma : f' a ≤ m) (hmb : m ≤ f' b) : + m ∈ f' '' Icc a b := +(ord_connected_Icc.image_has_deriv_within_at hf).out + (mem_image_of_mem _ (left_mem_Icc.2 hab)) (mem_image_of_mem _ (right_mem_Icc.2 hab)) ⟨hma, hmb⟩ + +/-- **Darboux's theorem**: if `a ≤ b` and `f' b ≤ m ≤ f' a`, then `f' c = m` for some +`c ∈ [a, b]`. -/ +theorem exists_has_deriv_within_at_eq_of_le_of_ge + (hab : a ≤ b) (hf : ∀ x ∈ (Icc a b), has_deriv_within_at f (f' x) (Icc a b) x) + {m : ℝ} (hma : f' a ≤ m) (hmb : m ≤ f' b) : + m ∈ f' '' Icc a b := +(ord_connected_Icc.image_has_deriv_within_at hf).out + (mem_image_of_mem _ (left_mem_Icc.2 hab)) (mem_image_of_mem _ (right_mem_Icc.2 hab)) ⟨hma, hmb⟩ + /-- If the derivative of a function is never equal to `m`, then either it is always greater than `m`, or it is always less than `m`. -/ -theorem deriv_forall_lt_or_forall_gt_of_forall_ne {s : set ℝ} (hs : convex ℝ s) - (hf : ∀ x ∈ s, has_deriv_at f (f' x) x) {m : ℝ} (hf' : ∀ x ∈ s, f' x ≠ m) : +theorem has_deriv_within_at_forall_lt_or_forall_gt_of_forall_ne {s : set ℝ} (hs : convex ℝ s) + (hf : ∀ x ∈ s, has_deriv_within_at f (f' x) s x) {m : ℝ} (hf' : ∀ x ∈ s, f' x ≠ m) : (∀ x ∈ s, f' x < m) ∨ (∀ x ∈ s, m < f' x) := begin contrapose! hf', rcases hf' with ⟨⟨b, hb, hmb⟩, ⟨a, ha, hma⟩⟩, - exact (convex_image_has_deriv_at hs hf).ord_connected.out (mem_image_of_mem f' ha) + exact (hs.ord_connected.image_has_deriv_within_at hf).out (mem_image_of_mem f' ha) (mem_image_of_mem f' hb) ⟨hma, hmb⟩ end From 74c2af38a828107941029b03839882c5c6f87a04 Mon Sep 17 00:00:00 2001 From: Peter Nelson Date: Tue, 30 May 2023 14:04:02 +0000 Subject: [PATCH 1114/1291] feat(data/set/ncard): noncomputable set cardinality (#18576) This PR introduces a function `set.ncard` that (noncomputably) gives the cardinality of a finite `set`, or zero if the set is infinite. The intention is to give a way to reason about set cardinality that avoids data and going between `set` and `finset`. This is especially useful for reasoning about sets in a `finite` type. The added file is somewhat large, but it is essentially just duplicating the API in `data.finset.card` (minus the induction lemmas). Co-authored-by: Peter Nelson <71660771+apnelson1@users.noreply.github.com> --- src/data/set/ncard.lean | 744 ++++++++++++++++++++++++++++++++++++++++ 1 file changed, 744 insertions(+) create mode 100644 src/data/set/ncard.lean diff --git a/src/data/set/ncard.lean b/src/data/set/ncard.lean new file mode 100644 index 0000000000000..965ef766b1616 --- /dev/null +++ b/src/data/set/ncard.lean @@ -0,0 +1,744 @@ +/- +Copyright (c) 2023 Peter Nelson. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Peter Nelson +-/ +import data.finite.card +import algebra.big_operators.finprod + +/-! +# Noncomputable Set Cardinality + +We define the cardinality `set.ncard s` of a set `s` as a natural number. This function is +noncomputable (being defined in terms of `nat.card`) and takes the value `0` if `s` is infinite. + +This can be seen as an API for `nat.card α` in the special case where `α` is a subtype arising from +a set. It is intended as an alternative to `finset.card` and `fintype.card`, both of which contain +data in their definition that can cause awkwardness when using `set.to_finset`. Using `set.ncard` +allows cardinality computations to avoid `finset`/`fintype` completely, staying in `set` and letting +finiteness be handled explicitly, or (where a `finite α` instance is present and the sets are +in `set α`) via `auto_param`s. + +## Main Definitions + +* `set.ncard s` is the cardinality of the set `s` as a natural number, provided `s` is finite. + If `s` is infinite, then `set.ncard s = 0`. +* `to_finite_tac` is a tactic that tries to synthesize an `set.finite s` argument with + `set.to_finite`. This will work for `s : set α` where there is a `finite α` instance. + +## Implementation Notes + +The lemmas in this file are very similar to those in `data.finset.card`, but with `set` operations +instead of `finset`; most of the proofs invoke their `finset` analogues. Nearly all the lemmas +require finiteness of one or more of their arguments. We provide this assumption with an +`auto_param` argument of the form `(hs : s.finite . to_finite_tac)`, where `to_finite_tac` will find +a `finite s` term in the cases where `s` is a set in a `finite` type. + +Often, where there are two set arguments `s` and `t`, the finiteness of one follows from the other +in the context of the lemma, in which case we only include the ones that are needed, and derive the +other inside the proof. A few of the lemmas, such as `ncard_union_le` do not require finiteness +arguments; they are are true by coincidence due to junk values. +-/ + +open_locale big_operators + +variables {α β : Type*} {s t : set α} {a b x y : α} {f : α → β} + +namespace set + +/-- The cardinality of `s : set α`. Has the junk value `0` if `s` is infinite -/ +noncomputable def ncard (s : set α) := nat.card s + +/-- A tactic, for use in `auto_param`s, that finds a `t.finite` term for a set `t` + whose finiteness can be deduced from typeclasses (eg. in a `finite` type). -/ +meta def to_finite_tac : tactic unit := `[exact set.to_finite _] + +lemma ncard_def (s : set α) : s.ncard = nat.card s := rfl + +lemma ncard_eq_to_finset_card (s : set α) (hs : s.finite . to_finite_tac) : + s.ncard = hs.to_finset.card := +by rw [ncard_def, @nat.card_eq_fintype_card _ hs.fintype, + @finite.card_to_finset _ _ hs.fintype hs] + +lemma ncard_le_of_subset (hst : s ⊆ t) (ht : t.finite . to_finite_tac) : s.ncard ≤ t.ncard := +@finite.card_le_of_embedding _ _ (finite_coe_iff.mpr ht) (set.embedding_of_subset _ _ hst) + +lemma ncard_mono [finite α] : @monotone (set α) _ _ _ ncard := +λ _ _, ncard_le_of_subset + +@[simp] lemma ncard_eq_zero (hs : s.finite . to_finite_tac) : s.ncard = 0 ↔ s = ∅ := +by simp [ncard_def, @finite.card_eq_zero_iff _ hs.to_subtype] + +@[simp] lemma ncard_coe_finset (s : finset α) : (s : set α).ncard = s.card := +by rw [ncard_eq_to_finset_card, finset.finite_to_set_to_finset] + +lemma infinite.ncard (hs : s.infinite) : s.ncard = 0 := +@nat.card_eq_zero_of_infinite _ hs.to_subtype + +lemma ncard_univ (α : Type*) : (univ : set α).ncard = nat.card α := +begin + cases finite_or_infinite α with h h, + { haveI := @fintype.of_finite α h, + rw [ncard_eq_to_finset_card, finite.to_finset_univ, finset.card_univ, + nat.card_eq_fintype_card]}, + rw [(@infinite_univ _ h).ncard, @nat.card_eq_zero_of_infinite _ h], +end + +@[simp] lemma ncard_empty (α : Type*) : (∅ : set α).ncard = 0 := +by simp only [ncard_eq_zero] + +lemma ncard_pos (hs : s.finite . to_finite_tac) : 0 < s.ncard ↔ s.nonempty := +by rw [pos_iff_ne_zero, ne.def, ncard_eq_zero hs, nonempty_iff_ne_empty] + +lemma ncard_ne_zero_of_mem (h : a ∈ s) (hs : s.finite . to_finite_tac) : s.ncard ≠ 0 := +((ncard_pos hs).mpr ⟨a,h⟩).ne.symm + +lemma finite_of_ncard_ne_zero (hs : s.ncard ≠ 0) : s.finite := +s.finite_or_infinite.elim id (λ h, (hs h.ncard).elim) + +lemma finite_of_ncard_pos (hs : 0 < s.ncard) : s.finite := +finite_of_ncard_ne_zero hs.ne.symm + +lemma nonempty_of_ncard_ne_zero (hs : s.ncard ≠ 0) : s.nonempty := +by { rw nonempty_iff_ne_empty, rintro rfl, simpa using hs } + +@[simp] lemma ncard_singleton (a : α) : ({a} : set α).ncard = 1 := +by simp [ncard_eq_to_finset_card] + +lemma ncard_singleton_inter : ({a} ∩ s).ncard ≤ 1 := +begin + classical, + rw [←inter_self {a}, inter_assoc, ncard_eq_to_finset_card, + finite.to_finset_inter, finite.to_finset_singleton], + { apply finset.card_singleton_inter}, + all_goals {apply to_finite}, +end + +section insert_erase + +@[simp] lemma ncard_insert_of_not_mem (h : a ∉ s) (hs : s.finite . to_finite_tac) : + (insert a s).ncard = s.ncard + 1 := +begin + classical, + haveI := hs.fintype, + rw [ncard_eq_to_finset_card, ncard_eq_to_finset_card, finite.to_finset_insert, + finset.card_insert_of_not_mem], + rwa [finite.mem_to_finset], +end + +lemma ncard_insert_of_mem (h : a ∈ s) : ncard (insert a s) = s.ncard := +by rw insert_eq_of_mem h + +lemma ncard_insert_le (a : α) (s : set α) : (insert a s).ncard ≤ s.ncard + 1 := +begin + classical, + obtain (hs | hs) := s.finite_or_infinite, + { exact (em (a ∈ s)).elim (λ h, (ncard_insert_of_mem h).trans_le (nat.le_succ _)) + (λ h, by rw ncard_insert_of_not_mem h hs)}, + rw (hs.mono (subset_insert a s)).ncard, + exact nat.zero_le _, +end + +lemma ncard_insert_eq_ite [decidable (a ∈ s)] (hs : s.finite . to_finite_tac) : + ncard (insert a s) = if a ∈ s then s.ncard else s.ncard + 1 := +begin + by_cases h : a ∈ s, + { rw [ncard_insert_of_mem h, if_pos h] }, + { rw [ncard_insert_of_not_mem h hs, if_neg h] } +end + +lemma ncard_le_ncard_insert (a : α) (s : set α) : s.ncard ≤ (insert a s).ncard := +begin + classical, + refine s.finite_or_infinite.elim (λ h, _) (λ h, by { rw h.ncard, exact nat.zero_le _ }), + rw ncard_insert_eq_ite h, + split_ifs, + { refl }, + { simp only [le_add_iff_nonneg_right, zero_le'] }, + exact classical.dec (a ∈ s), +end + +lemma ncard_pair (h : a ≠ b) : ({a, b} : set α).ncard = 2 := +by {rw [ncard_insert_of_not_mem, ncard_singleton], simpa} + +lemma ncard_diff_singleton_add_one (h : a ∈ s) (hs : s.finite . to_finite_tac) : + (s \ {a}).ncard + 1 = s.ncard := +begin + have h' : a ∉ s \ {a}, by {rw [mem_diff_singleton], tauto}, + rw ←ncard_insert_of_not_mem h' (hs.diff {a}), + congr', + simpa, +end + +lemma ncard_diff_singleton_of_mem (h : a ∈ s) (hs : s.finite . to_finite_tac) : + (s \ {a}).ncard = s.ncard - 1 := +eq_tsub_of_add_eq (ncard_diff_singleton_add_one h hs) + +lemma ncard_diff_singleton_lt_of_mem (h : a ∈ s) (hs : s.finite . to_finite_tac) : + (s \ {a}).ncard < s.ncard := +by {rw [←ncard_diff_singleton_add_one h hs], apply lt_add_one} + +lemma ncard_diff_singleton_le (s : set α) (a : α) : (s \ {a}).ncard ≤ s.ncard := +begin + obtain (hs | hs) := s.finite_or_infinite, + { apply ncard_le_of_subset (diff_subset _ _) hs}, + convert zero_le _, + exact (hs.diff (by simp : set.finite {a})).ncard, +end + +lemma pred_ncard_le_ncard_diff_singleton (s : set α) (a : α) : s.ncard - 1 ≤ (s \ {a}).ncard := +begin + cases s.finite_or_infinite with hs hs, + { by_cases h : a ∈ s, + { rw ncard_diff_singleton_of_mem h hs, }, + rw diff_singleton_eq_self h, + apply nat.pred_le}, + convert nat.zero_le _, + rw hs.ncard, +end + +lemma ncard_exchange (ha : a ∉ s) (hb : b ∈ s) : (insert a (s \ {b})).ncard = s.ncard := +begin + cases s.finite_or_infinite with h h, + { haveI := h.to_subtype, + rw [ncard_insert_of_not_mem, ncard_diff_singleton_add_one hb], + simpa only [mem_diff, not_and] using ha}, + rw [((h.diff (set.to_finite {b})).mono (subset_insert _ _)).ncard, h.ncard], +end + +lemma ncard_exchange' (ha : a ∉ s) (hb : b ∈ s) : ((insert a s) \ {b}).ncard = s.ncard := +by rw [←ncard_exchange ha hb, ←singleton_union, ←singleton_union, union_diff_distrib, + @diff_singleton_eq_self _ b {a} (λ h, ha (by rwa ← mem_singleton_iff.mp h) )] + +end insert_erase + +lemma ncard_image_le (hs : s.finite . to_finite_tac) : (f '' s).ncard ≤ s.ncard := +begin + classical, + rw ncard_eq_to_finset_card s hs, + haveI := hs.fintype, + convert @finset.card_image_le _ _ s.to_finset f _, + rw [ncard_eq_to_finset_card, finite.to_finset_image _ hs], + { congr', rw [←finset.coe_inj, finite.coe_to_finset, coe_to_finset]}, + { apply_instance}, + rw [←finset.coe_inj, finite.coe_to_finset, coe_to_finset], +end + +lemma ncard_image_of_inj_on (H : set.inj_on f s) : (f '' s).ncard = s.ncard := +begin + cases s.finite_or_infinite, + { haveI := @fintype.of_finite s h.to_subtype, + haveI := @fintype.of_finite _ (h.image f).to_subtype, + convert card_image_of_inj_on H; simp [ncard_def]}, + rw [h.ncard, ((infinite_image_iff H).mpr h).ncard], +end + +lemma inj_on_of_ncard_image_eq (h : (f '' s).ncard = s.ncard) (hs : s.finite . to_finite_tac) : + set.inj_on f s := +begin + classical, + haveI := hs.fintype, + haveI := ((to_finite s).image f).fintype, + simp_rw ncard_eq_to_finset_card at h, + rw ← coe_to_finset s, + apply finset.inj_on_of_card_image_eq, + convert h, + ext, + simp, +end + +lemma ncard_image_iff (hs : s.finite . to_finite_tac) : (f '' s).ncard = s.ncard ↔ set.inj_on f s := +⟨λ h, inj_on_of_ncard_image_eq h hs, ncard_image_of_inj_on⟩ + +lemma ncard_image_of_injective (s : set α) (H : f.injective) : (f '' s).ncard = s.ncard := +ncard_image_of_inj_on $ λ x _ y _ h, H h + +lemma ncard_preimage_of_injective_subset_range {s : set β} (H : f.injective) + (hs : s ⊆ set.range f) : (f ⁻¹' s).ncard = s.ncard := +by rw [←ncard_image_of_injective _ H, image_preimage_eq_iff.mpr hs] + +lemma fiber_ncard_ne_zero_iff_mem_image {y : β} (hs : s.finite . to_finite_tac) : + {x ∈ s | f x = y}.ncard ≠ 0 ↔ y ∈ f '' s := +begin + refine ⟨nonempty_of_ncard_ne_zero, _⟩, + rintros ⟨z,hz,rfl⟩, + exact @ncard_ne_zero_of_mem _ {x ∈ s | f x = f z} z (mem_sep hz rfl) + (hs.subset (sep_subset _ _)), +end + +@[simp] lemma ncard_map (f : α ↪ β) : (f '' s).ncard = s.ncard := +ncard_image_of_injective _ f.injective + +@[simp] lemma ncard_subtype (P : α → Prop) (s : set α) : + {x : subtype P | (x : α) ∈ s}.ncard = (s ∩ (set_of P)).ncard := +begin + convert (ncard_image_of_injective _ (@subtype.coe_injective _ P)).symm, + ext, rw inter_comm, simp, +end + +@[simp] lemma nat.card_coe_set_eq (s : set α) : nat.card s = s.ncard := +begin + convert (ncard_image_of_injective univ subtype.coe_injective).symm using 1, + { rw ncard_univ, refl }, + simp, +end + +lemma ncard_inter_le_ncard_left (s t : set α) (hs : s.finite . to_finite_tac) : + (s ∩ t).ncard ≤ s.ncard := +ncard_le_of_subset (inter_subset_left _ _) hs + +lemma ncard_inter_le_ncard_right (s t : set α) (ht : t.finite . to_finite_tac) : + (s ∩ t).ncard ≤ t.ncard := +ncard_le_of_subset (inter_subset_right _ _) ht + +lemma eq_of_subset_of_ncard_le (h : s ⊆ t) (h' : t.ncard ≤ s.ncard) + (ht : t.finite . to_finite_tac) : + s = t := +begin + haveI := ht.fintype, + haveI := (ht.subset h).fintype, + rw ←@to_finset_inj, + apply finset.eq_of_subset_of_card_le, + { simpa, }, + rw [ncard_eq_to_finset_card _ ht, ncard_eq_to_finset_card _ (ht.subset h)] at h', + convert h', +end + +lemma subset_iff_eq_of_ncard_le (h : t.ncard ≤ s.ncard) (ht : t.finite . to_finite_tac) : + s ⊆ t ↔ s = t := +⟨λ hst, eq_of_subset_of_ncard_le hst h ht, eq.subset'⟩ + +lemma map_eq_of_subset {f : α ↪ α} (h : f '' s ⊆ s) (hs : s.finite . to_finite_tac) : f '' s = s := +eq_of_subset_of_ncard_le h (ncard_map _).ge hs + +lemma sep_of_ncard_eq {P : α → Prop} (h : {x ∈ s | P x}.ncard = s.ncard) (ha : a ∈ s) + (hs : s.finite . to_finite_tac) : +P a := +sep_eq_self_iff_mem_true.mp (eq_of_subset_of_ncard_le (by simp) h.symm.le hs) _ ha + +lemma ncard_lt_ncard (h : s ⊂ t) (ht : t.finite . to_finite_tac) : s.ncard < t.ncard := +begin + rw [ncard_eq_to_finset_card _ (ht.subset h.subset), ncard_eq_to_finset_card t ht], + refine finset.card_lt_card _, + rwa [finite.to_finset_ssubset_to_finset], +end + +lemma ncard_strict_mono [finite α] : @strict_mono (set α) _ _ _ ncard := +λ _ _ h, ncard_lt_ncard h + +lemma ncard_eq_of_bijective {n : ℕ} (f : ∀ i, i < n → α) (hf : ∀ a ∈ s, ∃ i, ∃ h : i < n, f i h = a) + (hf' : ∀ i (h : i < n), f i h ∈ s) + (f_inj : ∀ i j (hi : i < n) (hj : j < n), f i hi = f j hj → i = j) + (hs : s.finite . to_finite_tac) : + s.ncard = n := +begin + rw ncard_eq_to_finset_card _ hs, + apply finset.card_eq_of_bijective, + all_goals {simpa}, +end + +lemma ncard_congr {t : set β} (f : Π a ∈ s, β) (h₁ : ∀ a ha, f a ha ∈ t) + (h₂ : ∀ a b ha hb, f a ha = f b hb → a = b) (h₃ : ∀ b ∈ t, ∃ a ha, f a ha = b) + (hs : s.finite . to_finite_tac) : + s.ncard = t.ncard := +begin + set f' : s → t := λ x, ⟨f x.1 x.2, h₁ _ _⟩ with hf', + have hbij : f'.bijective, + { split, + { rintros ⟨x,hx⟩ ⟨y,hy⟩ hxy, + simp only [hf', subtype.val_eq_coe, subtype.coe_mk, subtype.mk_eq_mk] at hxy ⊢, + apply h₂ _ _ hx hy hxy}, + rintro ⟨y,hy⟩, + obtain ⟨a, ha, rfl⟩ := h₃ y hy, + simp only [subtype.val_eq_coe, subtype.coe_mk, subtype.mk_eq_mk, set_coe.exists], + exact ⟨_,ha,rfl⟩}, + haveI := hs.to_subtype, + haveI := @fintype.of_finite _ (finite.of_bijective hbij), + haveI := fintype.of_finite s, + convert fintype.card_of_bijective hbij, + rw [ncard_def, nat.card_eq_fintype_card], + rw [ncard_def, nat.card_eq_fintype_card], +end + +lemma ncard_le_ncard_of_inj_on {t : set β} (f : α → β) (hf : ∀ a ∈ s, f a ∈ t) (f_inj : inj_on f s) + (ht : t.finite . to_finite_tac) : + s.ncard ≤ t.ncard := +begin + cases s.finite_or_infinite, + { haveI := h.to_subtype, + rw [ncard_eq_to_finset_card _ ht, ncard_eq_to_finset_card _ (to_finite s)], + exact finset.card_le_card_of_inj_on f (by simpa) (by simpa)}, + convert nat.zero_le _, + rw h.ncard, +end + +lemma exists_ne_map_eq_of_ncard_lt_of_maps_to {t : set β} (hc : t.ncard < s.ncard) {f : α → β} + (hf : ∀ a ∈ s, f a ∈ t) (ht : t.finite . to_finite_tac) : + ∃ (x ∈ s) (y ∈ s), x ≠ y ∧ f x = f y := +begin + by_contra h', + simp only [ne.def, exists_prop, not_exists, not_and, not_imp_not] at h', + exact (ncard_le_ncard_of_inj_on f hf h' ht).not_lt hc, +end + +lemma le_ncard_of_inj_on_range {n : ℕ} (f : ℕ → α) (hf : ∀ i < n, f i ∈ s) + (f_inj : ∀ (i < n) (j < n), f i = f j → i = j) (hs : s.finite . to_finite_tac): + n ≤ s.ncard := +by {rw ncard_eq_to_finset_card _ hs, apply finset.le_card_of_inj_on_range; simpa} + +lemma surj_on_of_inj_on_of_ncard_le {t : set β} (f : Π a ∈ s, β) + (hf : ∀ a ha, f a ha ∈ t) (hinj : ∀ a₁ a₂ ha₁ ha₂, f a₁ ha₁ = f a₂ ha₂ → a₁ = a₂) + (hst : t.ncard ≤ s.ncard) (ht : t.finite . to_finite_tac) : + ∀ b ∈ t, ∃ a ha, b = f a ha := +begin + intros b hb, + set f' : s → t := λ x, ⟨f x.1 x.2, hf _ _⟩ with hf', + have finj: f'.injective, + { rintros ⟨x,hx⟩ ⟨y,hy⟩ hxy, + simp only [hf', subtype.val_eq_coe, subtype.coe_mk, subtype.mk_eq_mk] at hxy ⊢, + apply hinj _ _ hx hy hxy}, + haveI := ht.fintype, + haveI := fintype.of_injective f' finj, + simp_rw [ncard_eq_to_finset_card] at hst, + set f'' : ∀ a, a ∈ s.to_finset → β := λ a h, f a (by simpa using h) with hf'', + convert @finset.surj_on_of_inj_on_of_card_le _ _ _ t.to_finset f'' (by simpa) (by simpa) + (by convert hst) b (by simpa), + simp, +end + +lemma inj_on_of_surj_on_of_ncard_le {t : set β} (f : Π a ∈ s, β) + (hf : ∀ a ha, f a ha ∈ t) (hsurj : ∀ b ∈ t, ∃ a ha, b = f a ha) (hst : s.ncard ≤ t.ncard) + ⦃a₁ a₂⦄ (ha₁ : a₁ ∈ s) (ha₂ : a₂ ∈ s) (ha₁a₂: f a₁ ha₁ = f a₂ ha₂) + (hs : s.finite . to_finite_tac) : + a₁ = a₂ := +begin + classical, + set f' : s → t := λ x, ⟨f x.1 x.2, hf _ _⟩ with hf', + have hsurj : f'.surjective, + { rintro ⟨y,hy⟩, + obtain ⟨a, ha, rfl⟩ := hsurj y hy, + simp only [subtype.val_eq_coe, subtype.coe_mk, subtype.mk_eq_mk, set_coe.exists], + exact ⟨_,ha,rfl⟩}, + haveI := hs.fintype, + haveI := fintype.of_surjective _ hsurj, + simp_rw [ncard_eq_to_finset_card] at hst, + set f'' : ∀ a, a ∈ s.to_finset → β := λ a h, f a (by simpa using h) with hf'', + exact @finset.inj_on_of_surj_on_of_card_le _ _ _ t.to_finset f'' (by simpa) (by simpa) + (by convert hst) a₁ a₂ (by simpa) (by simpa) (by simpa), +end + +section lattice + +lemma ncard_union_add_ncard_inter (s t : set α) (hs : s.finite . to_finite_tac) + (ht : t.finite . to_finite_tac) : + (s ∪ t).ncard + (s ∩ t).ncard = s.ncard + t.ncard := +begin + classical, + have hu := hs.union ht, + have hi := (hs.subset (inter_subset_left s t)), + rw [ncard_eq_to_finset_card _ hs, ncard_eq_to_finset_card _ ht, ncard_eq_to_finset_card _ hu, + ncard_eq_to_finset_card _ hi, finite.to_finset_union, finite.to_finset_inter], + { exact finset.card_union_add_card_inter _ _}, +end + +lemma ncard_inter_add_ncard_union (s t : set α) (hs : s.finite . to_finite_tac) + (ht : t.finite . to_finite_tac) : + (s ∩ t).ncard + (s ∪ t).ncard = s.ncard + t.ncard := +by rw [add_comm, ncard_union_add_ncard_inter _ _ hs ht] + +lemma ncard_union_le (s t : set α) : (s ∪ t).ncard ≤ s.ncard + t.ncard := +begin + classical, + cases (s ∪ t).finite_or_infinite, + { have hs := h.subset (subset_union_left s t), + have ht := h.subset (subset_union_right s t), + rw [ncard_eq_to_finset_card _ hs, ncard_eq_to_finset_card _ ht, ncard_eq_to_finset_card _ h, + finite.to_finset_union], + exact finset.card_union_le _ _}, + convert nat.zero_le _, + rw h.ncard, +end + +lemma ncard_union_eq (h : disjoint s t) (hs : s.finite . to_finite_tac) + (ht : t.finite . to_finite_tac) : + (s ∪ t).ncard = s.ncard + t.ncard := +begin + classical, + rw [ncard_eq_to_finset_card _ hs, ncard_eq_to_finset_card _ ht, + ncard_eq_to_finset_card _ (hs.union ht),finite.to_finset_union], + refine finset.card_union_eq _, + rwa [finite.disjoint_to_finset], +end + +lemma ncard_diff_add_ncard_eq_ncard (h : s ⊆ t) (ht : t.finite . to_finite_tac) : + (t \ s).ncard + s.ncard = t.ncard := +begin + classical, + rw [ncard_eq_to_finset_card _ ht, ncard_eq_to_finset_card _ (ht.subset h), + ncard_eq_to_finset_card _ (ht.diff s), finite.to_finset_diff], + refine finset.card_sdiff_add_card_eq_card _, + rwa finite.to_finset_subset_to_finset, +end + +lemma ncard_diff (h : s ⊆ t) (ht : t.finite . to_finite_tac) : (t \ s).ncard = t.ncard - s.ncard := +by rw [←ncard_diff_add_ncard_eq_ncard h ht, add_tsub_cancel_right] + +lemma ncard_le_ncard_diff_add_ncard (s t : set α) (ht : t.finite . to_finite_tac) : + s.ncard ≤ (s \ t).ncard + t.ncard := +begin + cases s.finite_or_infinite, + { rw [←diff_inter_self_eq_diff, ←ncard_diff_add_ncard_eq_ncard (inter_subset_right t s) h, + add_le_add_iff_left], + apply ncard_inter_le_ncard_left _ _ ht,}, + convert nat.zero_le _, + rw h.ncard, +end + +lemma le_ncard_diff (s t : set α) (hs : s.finite . to_finite_tac) : + t.ncard - s.ncard ≤ (t \ s).ncard := +begin + refine tsub_le_iff_left.mpr _, + rw add_comm, + apply ncard_le_ncard_diff_add_ncard _ _ hs, +end + +lemma ncard_diff_add_ncard (s t : set α) (hs : s.finite . to_finite_tac) + (ht : t.finite . to_finite_tac) : + (s \ t).ncard + t.ncard = (s ∪ t).ncard := +by rw [←union_diff_right,ncard_diff_add_ncard_eq_ncard (subset_union_right s t) (hs.union ht)] + +lemma diff_nonempty_of_ncard_lt_ncard (h : s.ncard < t.ncard) (hs : s.finite . to_finite_tac) : + (t \ s).nonempty := +begin + rw [set.nonempty_iff_ne_empty, ne.def, diff_eq_empty], + exact λ h', h.not_le (ncard_le_of_subset h' hs), +end + +lemma exists_mem_not_mem_of_ncard_lt_ncard (h : s.ncard < t.ncard) (hs : s.finite . to_finite_tac) : + ∃ e, e ∈ t ∧ e ∉ s := +diff_nonempty_of_ncard_lt_ncard h hs + +@[simp] lemma ncard_inter_add_ncard_diff_eq_ncard (s t : set α) + (hs : s.finite . to_finite_tac) : + (s ∩ t).ncard + (s \ t).ncard = s.ncard := +by rw [←ncard_diff_add_ncard_eq_ncard (diff_subset s t) hs, sdiff_sdiff_right_self, inf_eq_inter] + +lemma ncard_eq_ncard_iff_ncard_diff_eq_ncard_diff (hs : s.finite . to_finite_tac) + (ht : t.finite . to_finite_tac) : + s.ncard = t.ncard ↔ (s \ t).ncard = (t \ s).ncard := +by rw [←ncard_inter_add_ncard_diff_eq_ncard s t hs, ←ncard_inter_add_ncard_diff_eq_ncard t s ht, + inter_comm, add_right_inj] + +lemma ncard_le_ncard_iff_ncard_diff_le_ncard_diff (hs : s.finite . to_finite_tac) + (ht : t.finite . to_finite_tac) : + s.ncard ≤ t.ncard ↔ (s \ t).ncard ≤ (t \ s).ncard := +by rw [←ncard_inter_add_ncard_diff_eq_ncard s t hs, ←ncard_inter_add_ncard_diff_eq_ncard t s ht, + inter_comm, add_le_add_iff_left] + +lemma ncard_lt_ncard_iff_ncard_diff_lt_ncard_diff (hs : s.finite . to_finite_tac) + (ht : t.finite . to_finite_tac) : + s.ncard < t.ncard ↔ (s \ t).ncard < (t \ s).ncard := +by rw [←ncard_inter_add_ncard_diff_eq_ncard s t hs, ←ncard_inter_add_ncard_diff_eq_ncard t s ht, + inter_comm, add_lt_add_iff_left] + +lemma ncard_add_ncard_compl (s : set α) (hs : s.finite . to_finite_tac) + (hsc : sᶜ.finite . to_finite_tac) : + s.ncard + sᶜ.ncard = nat.card α := +by rw [←ncard_univ, ←ncard_union_eq (@disjoint_compl_right _ _ s) hs hsc, union_compl_self] + +end lattice + +/-- Given a set `t` and a set `s` inside it, we can shrink `t` to any appropriate size, and keep `s` + inside it. -/ +lemma exists_intermediate_set (i : ℕ) (h₁ : i + s.ncard ≤ t.ncard) (h₂ : s ⊆ t) : + ∃ (r : set α), s ⊆ r ∧ r ⊆ t ∧ r.ncard = i + s.ncard := +begin + cases t.finite_or_infinite with ht ht, + { haveI := ht.to_subtype, + haveI := (ht.subset h₂).to_subtype, + simp_rw [ncard_eq_to_finset_card] at h₁ ⊢, + obtain ⟨r', hsr', hr't, hr'⟩ := finset.exists_intermediate_set _ h₁ (by simpa), + exact ⟨r', by simpa using hsr', by simpa using hr't, by rw [←hr', ncard_coe_finset]⟩}, + rw [ht.ncard] at h₁, + have h₁' := nat.eq_zero_of_le_zero h₁, + rw [add_eq_zero_iff] at h₁', + exact ⟨t, h₂, rfl.subset, by rw [ht.ncard, h₁'.1, h₁'.2]⟩ +end + +lemma exists_intermediate_set' {m : ℕ} (hs : s.ncard ≤ m) (ht : m ≤ t.ncard) (h : s ⊆ t) : + ∃ (r : set α), s ⊆ r ∧ r ⊆ t ∧ r.ncard = m := +begin + obtain ⟨r,hsr,hrt,hc⟩ := + exists_intermediate_set (m - s.ncard) (by rwa [tsub_add_cancel_of_le hs]) h, + rw tsub_add_cancel_of_le hs at hc, + exact ⟨r,hsr,hrt,hc⟩, +end + +/-- We can shrink `s` to any smaller size. -/ +lemma exists_smaller_set (s : set α) (i : ℕ) (h₁ : i ≤ s.ncard) : + ∃ (t : set α), t ⊆ s ∧ t.ncard = i := +(exists_intermediate_set i (by simpa) (empty_subset s)).imp (λ t ht, ⟨ht.2.1,by simpa using ht.2.2⟩) + +lemma infinite.exists_subset_ncard_eq {s : set α} (hs : s.infinite) (k : ℕ) : + ∃ t, t ⊆ s ∧ t.finite ∧ t.ncard = k := +begin + haveI := hs.to_subtype, + obtain ⟨t', -, rfl⟩ := @infinite.exists_subset_card_eq s univ infinite_univ k, + refine ⟨coe '' (t' : set s), by simp, finite.image _ (by simp), _⟩, + rw [ncard_image_of_injective _ subtype.coe_injective], + simp, +end + +lemma infinite.exists_supset_ncard_eq {s t : set α} (ht : t.infinite) (hst : s ⊆ t) + (hs : s.finite) {k : ℕ} (hsk : s.ncard ≤ k) : + ∃ s', s ⊆ s' ∧ s' ⊆ t ∧ s'.ncard = k := +begin + obtain ⟨s₁, hs₁, hs₁fin, hs₁card⟩ := (ht.diff hs).exists_subset_ncard_eq (k - s.ncard), + refine ⟨s ∪ s₁, subset_union_left _ _, union_subset hst (hs₁.trans (diff_subset _ _)), _⟩, + rwa [ncard_union_eq (disjoint_of_subset_right hs₁ disjoint_sdiff_right) hs hs₁fin, hs₁card, + add_tsub_cancel_of_le], +end + +lemma exists_subset_or_subset_of_two_mul_lt_ncard {n : ℕ} (hst : 2 * n < (s ∪ t).ncard) : + ∃ (r : set α), n < r.ncard ∧ (r ⊆ s ∨ r ⊆ t) := +begin + classical, + have hu := (finite_of_ncard_ne_zero ((nat.zero_le _).trans_lt hst).ne.symm), + rw [ncard_eq_to_finset_card _ hu, finite.to_finset_union + (hu.subset (subset_union_left _ _)) (hu.subset (subset_union_right _ _))] at hst, + obtain ⟨r', hnr', hr'⟩ := finset.exists_subset_or_subset_of_two_mul_lt_card hst, + exact ⟨r', by simpa , by simpa using hr'⟩, +end + +/-! ### Explicit description of a set from its cardinality -/ + +@[simp] lemma ncard_eq_one : s.ncard = 1 ↔ ∃ a, s = {a} := +begin + refine ⟨λ h, _,by {rintro ⟨a,rfl⟩, rw [ncard_singleton]}⟩, + haveI := (finite_of_ncard_ne_zero (ne_zero_of_eq_one h)).to_subtype, + rw [ncard_eq_to_finset_card, finset.card_eq_one] at h, + exact h.imp (λ a ha, by rwa [←finite.to_finset_singleton, finite.to_finset_inj] at ha), +end + +lemma exists_eq_insert_iff_ncard (hs : s.finite . to_finite_tac) : + (∃ a ∉ s, insert a s = t) ↔ s ⊆ t ∧ s.ncard + 1 = t.ncard := +begin + classical, + split, + { rintro ⟨a, ha, rfl⟩, + rw [ncard_eq_to_finset_card _ hs, ncard_eq_to_finset_card _ (hs.insert a), + finite.to_finset_insert, ←@finite.to_finset_subset_to_finset _ _ _ hs (hs.insert a), + finite.to_finset_insert], + refine (@finset.exists_eq_insert_iff _ _ hs.to_finset (insert a hs.to_finset)).mp _, + exact ⟨a, by rwa finite.mem_to_finset, rfl⟩}, + rintro ⟨hst, h⟩, + have ht := @finite_of_ncard_pos _ t (by {rw ←h, apply nat.zero_lt_succ}), + + rw [ncard_eq_to_finset_card _ hs, ncard_eq_to_finset_card _ ht] at h, + obtain ⟨a,has, ha⟩ := (finset.exists_eq_insert_iff.mpr ⟨by {simpa},h⟩), + have hsa := hs.insert a, + rw ←finite.to_finset_insert at ha, + exact ⟨a, by {rwa finite.mem_to_finset at has}, by {rwa ←@finite.to_finset_inj _ _ _ hsa ht}⟩, +end + +lemma ncard_le_one (hs : s.finite . to_finite_tac) : s.ncard ≤ 1 ↔ ∀ (a ∈ s) (b ∈ s), a = b := +by simp_rw [ncard_eq_to_finset_card _ hs, finset.card_le_one, finite.mem_to_finset] + +lemma ncard_le_one_iff (hs : s.finite . to_finite_tac) : + s.ncard ≤ 1 ↔ ∀ {a b}, a ∈ s → b ∈ s → a = b := +by { rw ncard_le_one hs, tauto} + +lemma ncard_le_one_iff_eq (hs : s.finite . to_finite_tac) : s.ncard ≤ 1 ↔ s = ∅ ∨ ∃ a, s = {a} := +begin + obtain (rfl | ⟨x,hx⟩) := s.eq_empty_or_nonempty, + { exact iff_of_true (by simp) (or.inl rfl), }, + rw [ncard_le_one_iff hs], + refine ⟨λ h, or.inr ⟨x,(singleton_subset_iff.mpr hx).antisymm' (λ y hy, h hy hx)⟩, _⟩, + rintro (rfl | ⟨a,rfl⟩), + { exact (not_mem_empty _ hx).elim }, + simp_rw mem_singleton_iff at hx ⊢, subst hx, + exact λ a b h h', h.trans h'.symm, +end + +lemma ncard_le_one_iff_subset_singleton [nonempty α] (hs : s.finite . to_finite_tac) : + s.ncard ≤ 1 ↔ ∃ (x : α), s ⊆ {x} := +by simp_rw [ncard_eq_to_finset_card _ hs, finset.card_le_one_iff_subset_singleton, + finite.to_finset_subset, finset.coe_singleton] + +/-- A `set` of a subsingleton type has cardinality at most one. -/ +lemma ncard_le_one_of_subsingleton [subsingleton α] (s : set α) : s.ncard ≤ 1 := +by {rw [ncard_eq_to_finset_card], exact finset.card_le_one_of_subsingleton _} + +lemma one_lt_ncard (hs : s.finite . to_finite_tac) : 1 < s.ncard ↔ ∃ (a ∈ s) (b ∈ s), a ≠ b := +by simp_rw [ncard_eq_to_finset_card _ hs, finset.one_lt_card, finite.mem_to_finset] + +lemma one_lt_ncard_iff (hs : s.finite . to_finite_tac) : + 1 < s.ncard ↔ ∃ a b, a ∈ s ∧ b ∈ s ∧ a ≠ b := +by { rw one_lt_ncard hs, simp only [exists_prop, exists_and_distrib_left] } + +lemma two_lt_ncard_iff (hs : s.finite . to_finite_tac) : + 2 < s.ncard ↔ ∃ a b c, a ∈ s ∧ b ∈ s ∧ c ∈ s ∧ a ≠ b ∧ a ≠ c ∧ b ≠ c := +by simp_rw [ncard_eq_to_finset_card _ hs, finset.two_lt_card_iff, finite.mem_to_finset] + +lemma two_lt_card (hs : s.finite . to_finite_tac) : + 2 < s.ncard ↔ ∃ (a ∈ s) (b ∈ s) (c ∈ s), a ≠ b ∧ a ≠ c ∧ b ≠ c := +by simp only [two_lt_ncard_iff hs, exists_and_distrib_left, exists_prop] + +lemma exists_ne_of_one_lt_ncard (hs : 1 < s.ncard) (a : α) : ∃ b, b ∈ s ∧ b ≠ a := +begin + haveI := (finite_of_ncard_ne_zero (zero_lt_one.trans hs).ne.symm).to_subtype, + rw [ncard_eq_to_finset_card] at hs, + simpa only [finite.mem_to_finset] using finset.exists_ne_of_one_lt_card hs a, +end + +lemma eq_insert_of_ncard_eq_succ {n : ℕ} (h : s.ncard = n + 1) : + ∃ a t, a ∉ t ∧ insert a t = s ∧ t.ncard = n := +begin + classical, + haveI := @fintype.of_finite _ (finite_of_ncard_pos (n.zero_lt_succ.trans_eq h.symm)).to_subtype, + rw [ncard_eq_to_finset_card, finset.card_eq_succ] at h, + obtain ⟨a,t,hat,hts,rfl⟩ := h, + refine ⟨a,t,hat,_,by rw ncard_coe_finset⟩, + rw [←to_finset_inj], + convert hts, + simp only [to_finset_insert, finset.to_finset_coe], +end + +lemma ncard_eq_succ {n : ℕ} (hs : s.finite . to_finite_tac) : + s.ncard = n + 1 ↔ ∃ a t, a ∉ t ∧ insert a t = s ∧ t.ncard = n := +begin + classical, + refine ⟨eq_insert_of_ncard_eq_succ, _⟩, + rintro ⟨a,t,hat,h,rfl⟩, + rw [← h, ncard_insert_of_not_mem hat (hs.subset ((subset_insert a t).trans_eq h))] +end + +lemma ncard_eq_two : s.ncard = 2 ↔ ∃ x y, x ≠ y ∧ s = {x, y} := +begin + classical, + refine ⟨λ h, _, _⟩, + { obtain ⟨x,t,hxt,rfl,ht⟩ := eq_insert_of_ncard_eq_succ h, + obtain ⟨y,rfl⟩ := ncard_eq_one.mp ht, + rw mem_singleton_iff at hxt, + exact ⟨_,_,hxt,rfl⟩}, + rintro ⟨x,y,hxy,rfl⟩, + rw [ncard_eq_to_finset_card, finset.card_eq_two], + exact ⟨x,y,hxy, by {ext, simp}⟩, +end + +lemma ncard_eq_three : s.ncard = 3 ↔ ∃ x y z, x ≠ y ∧ x ≠ z ∧ y ≠ z ∧ s = {x, y, z} := +begin + classical, + refine ⟨λ h, _, _⟩, + { obtain ⟨x,t,hxt,rfl,ht⟩ := eq_insert_of_ncard_eq_succ h, + obtain ⟨y,z,hyz,rfl⟩ := ncard_eq_two.mp ht, + rw [mem_insert_iff, mem_singleton_iff, not_or_distrib] at hxt, + exact ⟨x,y,z,hxt.1,hxt.2,hyz,rfl⟩}, + rintro ⟨x, y, z, xy, xz, yz, rfl⟩, + rw [ncard_insert_of_not_mem, ncard_insert_of_not_mem, ncard_singleton], + { rwa mem_singleton_iff}, + rw [mem_insert_iff, mem_singleton_iff], + tauto, +end + +end set From 938d3db9c278f8a52c0f964a405806f0f2b09b74 Mon Sep 17 00:00:00 2001 From: jakelev Date: Tue, 30 May 2023 17:46:30 +0000 Subject: [PATCH 1115/1291] feat(ring_theory/polynomial/hermite): proof of explicit formula for Hermite polynomial coefficients (#19044) Add the explicit closed formula for the coefficients of Hermite polynomials. Co-authored-by: Luke Mantle --- src/ring_theory/polynomial/hermite/basic.lean | 81 ++++++++++++++++++- 1 file changed, 80 insertions(+), 1 deletion(-) diff --git a/src/ring_theory/polynomial/hermite/basic.lean b/src/ring_theory/polynomial/hermite/basic.lean index 4a6d951db090c..5df004058baac 100644 --- a/src/ring_theory/polynomial/hermite/basic.lean +++ b/src/ring_theory/polynomial/hermite/basic.lean @@ -6,7 +6,7 @@ Authors: Luke Mantle import data.polynomial.derivative import data.nat.parity - +import data.nat.factorial.double_factorial /-! # Hermite polynomials @@ -20,9 +20,14 @@ This file defines `polynomial.hermite n`, the nth probabilist's Hermite polynomi ## Results * `polynomial.hermite_succ`: the recursion `hermite (n+1) = (x - d/dx) (hermite n)` +* `polynomial.coeff_hermite_explicit`: a closed formula for (nonvanishing) coefficients in terms + of binomial coefficients and double factorials. * `polynomial.coeff_hermite_of_odd_add`: for `n`,`k` where `n+k` is odd, `(hermite n).coeff k` is zero. +* `polynomial.coeff_hermite_of_even_add`: a closed formula for `(hermite n).coeff k` when `n+k` is + even, equivalent to `polynomial.coeff_hermite_explicit`. * `polynomial.monic_hermite`: for all `n`, `hermite n` is monic. +* `polynomial.degree_hermite`: for all `n`, `hermite n` has degree `n`. ## References @@ -127,4 +132,78 @@ end end coeff +section coeff_explicit + +open_locale nat + +/-- Because of `coeff_hermite_of_odd_add`, every nonzero coefficient is described as follows. -/ +lemma coeff_hermite_explicit : + ∀ (n k : ℕ), coeff (hermite (2 * n + k)) k = (-1)^n * (2 * n - 1)‼ * nat.choose (2 * n + k) k +| 0 _ := by simp +| (n + 1) 0 := begin + convert coeff_hermite_succ_zero (2 * n + 1) using 1, + rw [coeff_hermite_explicit n 1, + (by ring_nf : 2 * (n + 1) - 1 = 2 * n + 1), nat.double_factorial_add_one, + nat.choose_zero_right, nat.choose_one_right, pow_succ], + push_cast, + ring, +end +| (n + 1) (k + 1) := begin + let hermite_explicit : ℕ → ℕ → ℤ := + λ n k, (-1)^n * (2*n-1)‼ * nat.choose (2*n+k) k, + have hermite_explicit_recur : + ∀ (n k : ℕ), + hermite_explicit (n + 1) (k + 1) = + hermite_explicit (n + 1) k - (k + 2) * hermite_explicit n (k + 2) := + begin + intros n k, + simp only [hermite_explicit], + -- Factor out (-1)'s. + rw [mul_comm (↑k + _), sub_eq_add_neg], + nth_rewrite 2 neg_eq_neg_one_mul, + simp only [mul_assoc, ← mul_add, pow_succ], + congr' 2, + -- Factor out double factorials. + norm_cast, + rw [(by ring_nf : 2 * (n + 1) - 1 = 2 * n + 1), + nat.double_factorial_add_one, mul_comm (2 * n + 1)], + simp only [mul_assoc, ← mul_add], + congr' 1, + -- Match up binomial coefficients using `nat.choose_succ_right_eq`. + rw [(by ring : 2 * (n + 1) + (k + 1) = (2 * n + 1) + (k + 1) + 1), + (by ring : 2 * (n + 1) + k = (2 * n + 1) + (k + 1)), + (by ring : 2 * n + (k + 2) = (2 * n + 1) + (k + 1))], + rw [nat.choose, nat.choose_succ_right_eq ((2 * n + 1) + (k + 1)) (k + 1), + nat.add_sub_cancel], + ring, + end, + change _ = hermite_explicit _ _, + rw [← add_assoc, coeff_hermite_succ_succ, hermite_explicit_recur], + congr, + { rw coeff_hermite_explicit (n + 1) k }, + { rw [(by ring : 2 * (n + 1) + k = 2 * n + (k + 2)), coeff_hermite_explicit n (k + 2)] }, +end + +lemma coeff_hermite_of_even_add {n k : ℕ} (hnk : even (n + k)) : + coeff (hermite n) k = (-1)^((n - k) / 2) * (n - k - 1)‼ * nat.choose n k := +begin + cases le_or_lt k n with h_le h_lt, + { rw [nat.even_add, ← (nat.even_sub h_le)] at hnk, + obtain ⟨m, hm⟩ := hnk, + rw [(by linarith : n = 2 * m + k), nat.add_sub_cancel, + nat.mul_div_cancel_left _ (nat.succ_pos 1), coeff_hermite_explicit] }, + { simp [nat.choose_eq_zero_of_lt h_lt, coeff_hermite_of_lt h_lt] }, +end + +lemma coeff_hermite (n k : ℕ) : + coeff (hermite n) k = + if even (n + k) then (-1)^((n - k) / 2) * (n - k - 1)‼ * nat.choose n k else 0 := +begin + split_ifs with h, + exact coeff_hermite_of_even_add h, + exact coeff_hermite_of_odd_add (nat.odd_iff_not_even.mpr h), +end + +end coeff_explicit + end polynomial From 48dc6abe71248bd6f4bffc9703dc87bdd4e37d0b Mon Sep 17 00:00:00 2001 From: MohanadAhmed Date: Tue, 30 May 2023 18:50:38 +0000 Subject: [PATCH 1116/1291] feat(linear_algebra/matrix/charpoly/eigs): det and trace are product and sum of eigenvalues (#19079) This adds two lemmas: 1. `det_eq_prod_roots_charpoly`: the determinant of a square matrix is the product of the characteristic polynomial roots of that matrix. 2. `trace_eq_sum_roots_charpoly`: the trace is similarly the sum of the characteristic polynomial roots. These two lemmas are more commonly stated as trace is the sum of eigenvalues and determinant is the product of eigenvalues. Mathlib has already defined eigenvalues in [linear_algebra.eigenspace](https://leanprover-community.github.io/mathlib_docs/linear_algebra/eigenspace.html#module.End.eigenvalues) as the roots of the minimal polynomial of a linear endomorphism. These do not have correct multiplicity and cannot be used in the theorems above. Hence we express these theorems in terms of the roots of the characteristic polynomial directly. Co-authored-by: Eric Wieser --- src/linear_algebra/matrix/charpoly/eigs.lean | 84 ++++++++++++++++++++ 1 file changed, 84 insertions(+) create mode 100644 src/linear_algebra/matrix/charpoly/eigs.lean diff --git a/src/linear_algebra/matrix/charpoly/eigs.lean b/src/linear_algebra/matrix/charpoly/eigs.lean new file mode 100644 index 0000000000000..a585f16e77072 --- /dev/null +++ b/src/linear_algebra/matrix/charpoly/eigs.lean @@ -0,0 +1,84 @@ +/- +Copyright (c) 2023 Mohanad Ahmed. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Mohanad Ahmed +-/ + +import data.polynomial.basic +import field_theory.is_alg_closed.basic + +/-! +# Eigenvalues are characteristic polynomial roots. + +In fields we show that: + +* `matrix.det_eq_prod_roots_charpoly_of_splits`: the determinant (in the field of the matrix) + is the product of the roots of the characteristic polynomial if the polynomial splits in the field + of the matrix. +* `matrix.trace_eq_sum_roots_charpoly_of_splits`: the trace is the sum of the roots of the + characteristic polynomial if the polynomial splits in the field of the matrix. + +In an algebraically closed field we show that: + +* `matrix.det_eq_prod_roots_charpoly`: the determinant is the product of the roots of the + characteristic polynomial. +* `matrix.trace_eq_sum_roots_charpoly`: the trace is the sum of the roots of the + characteristic polynomial. + +Note that over other fields such as `ℝ`, these results can be used by using +`A.map (algebra_map ℝ ℂ)` as the matrix, and then applying `ring_hom.map_det`. + +The two lemmas `matrix.det_eq_prod_roots_charpoly` and `matrix.trace_eq_sum_roots_charpoly` are more +commonly stated as trace is the sum of eigenvalues and determinant is the product of eigenvalues. +Mathlib has already defined eigenvalues in `linear_algebra.eigenspace` as the roots of the minimal +polynomial of a linear endomorphism. These do not have correct multiplicity and cannot be used in +the theorems above. Hence we express these theorems in terms of the roots of the characteristic +polynomial directly. + +## TODO + +The proofs of `det_eq_prod_roots_charpoly_of_splits` and +`trace_eq_sum_roots_charpoly_of_splits` closely resemble +`norm_gen_eq_prod_roots` and `trace_gen_eq_sum_roots` respectively, but the +dependencies are not general enough to unify them. We should refactor +`polynomial.prod_roots_eq_coeff_zero_of_monic_of_split` and +`polynomial.sum_roots_eq_next_coeff_of_monic_of_split` to assume splitting over an arbitrary map. +-/ +variables {n : Type*} [fintype n] [decidable_eq n] +variables {R : Type*} [field R] +variables {A : matrix n n R} + +open matrix polynomial +open_locale matrix big_operators + +namespace matrix + +lemma det_eq_prod_roots_charpoly_of_splits (hAps : A.charpoly.splits (ring_hom.id R)) : + A.det = (matrix.charpoly A).roots.prod := +begin + rw [det_eq_sign_charpoly_coeff, ← (charpoly_nat_degree_eq_dim A), + polynomial.prod_roots_eq_coeff_zero_of_monic_of_split A.charpoly_monic (hAps), + ← mul_assoc, ← pow_two, pow_right_comm, neg_one_sq, one_pow, one_mul], +end + +lemma trace_eq_sum_roots_charpoly_of_splits (hAps : A.charpoly.splits (ring_hom.id R)) : + A.trace = (matrix.charpoly A).roots.sum := +begin + casesI is_empty_or_nonempty n, + { rw [matrix.trace, fintype.sum_empty, matrix.charpoly, + det_eq_one_of_card_eq_zero (fintype.card_eq_zero_iff.2 h), polynomial.roots_one, + multiset.empty_eq_zero, multiset.sum_zero], }, + { rw [trace_eq_neg_charpoly_coeff, neg_eq_iff_eq_neg, + ← polynomial.sum_roots_eq_next_coeff_of_monic_of_split A.charpoly_monic (hAps), + next_coeff, charpoly_nat_degree_eq_dim, + if_neg (fintype.card_ne_zero : fintype.card n ≠ 0)], }, +end +variables (A) + +lemma det_eq_prod_roots_charpoly [is_alg_closed R] : A.det = (matrix.charpoly A).roots.prod := +det_eq_prod_roots_charpoly_of_splits (is_alg_closed.splits A.charpoly) + +lemma trace_eq_sum_roots_charpoly [is_alg_closed R] : A.trace = (matrix.charpoly A).roots.sum := +trace_eq_sum_roots_charpoly_of_splits (is_alg_closed.splits A.charpoly) + +end matrix From 328375597f2c0dd00522d9c2e5a33b6a6128feeb Mon Sep 17 00:00:00 2001 From: damiano Date: Tue, 30 May 2023 23:18:13 +0000 Subject: [PATCH 1117/1291] chore(archive + counterexamples): namespaces imo, theorems_100, counterexample, plus three more (#19129) This PR is a revision of #19122: it addresses namespacing in `archive` and `counterexamples`. The main difference with #19122 is that it adds namespaces less aggressively: I added a namespace only if there was not an explicit namespace after the initial "fluff". In `counterexamples`, I added namespaces to all files. I introduced three "main" namespaces: `imo, theorems_100, counterexample` (the last one singular). Besides these, I also introduced the namespaces `prop_encodable, oxford_invariants, sensitivity`, to cover the left-over files in `archive`. Note that if a file has `namespace` early on, then it does not get a new namespace, even though it might be desirable for it to have one. Comments are very welcome! Note: besides adding namespaces, the only files that I had to manually edit are the ones in * commit d337b99e3e6d147d440c91874d1809a3ee04ff16 -- I do not like these changes, but currently do not see how to avoid them; * commit 48471f35ec9f9929dd35363ee176f8c889042f6e -- I *removed* an pre-existing namespace, replacing it by `open `. [Zulip discussion](https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/mk_all.2Esh): this PR is motivated by the desire to port to mathlib4 the files in `archive, counterexamples`, taking advantage of `mathport`. The namespacing helps with avoiding clashes among names of declarations, as well as one-lettered declarations in the root namespace. --- .../37_solution_of_cubic.lean | 4 +++ .../42_inverse_triangle_sum.lean | 2 +- archive/100-theorems-list/45_partition.lean | 4 +++ .../100-theorems-list/57_herons_formula.lean | 4 +++ .../100-theorems-list/70_perfect_numbers.lean | 33 +++++++++++-------- .../73_ascending_descending_sequences.lean | 4 +++ .../81_sum_of_prime_reciprocals_diverges.lean | 4 +++ .../100-theorems-list/82_cubing_a_cube.lean | 4 +++ .../83_friendship_graphs.lean | 6 +++- .../93_birthday_problem.lean | 4 +++ .../100-theorems-list/9_area_of_a_circle.lean | 4 +++ archive/examples/prop_encodable.lean | 4 +++ archive/imo/imo1959_q1.lean | 4 +++ archive/imo/imo1962_q4.lean | 4 +++ archive/imo/imo1972_q5.lean | 4 +++ archive/imo/imo1975_q1.lean | 4 +++ archive/imo/imo1977_q6.lean | 4 +++ archive/imo/imo1988_q6.lean | 4 +++ archive/imo/imo1994_q1.lean | 4 +++ archive/imo/imo1998_q2.lean | 4 +++ archive/imo/imo2001_q2.lean | 4 +++ archive/imo/imo2001_q6.lean | 4 +++ archive/imo/imo2005_q3.lean | 4 +++ archive/imo/imo2005_q4.lean | 4 +++ archive/imo/imo2006_q3.lean | 4 +++ archive/imo/imo2006_q5.lean | 4 +++ archive/imo/imo2008_q2.lean | 4 +++ archive/imo/imo2008_q3.lean | 4 +++ archive/imo/imo2008_q4.lean | 4 +++ archive/imo/imo2011_q3.lean | 4 +++ archive/imo/imo2011_q5.lean | 4 +++ archive/imo/imo2013_q1.lean | 4 +++ archive/imo/imo2013_q5.lean | 4 +++ archive/imo/imo2019_q1.lean | 4 +++ archive/imo/imo2019_q2.lean | 4 +++ archive/imo/imo2019_q4.lean | 4 +++ archive/imo/imo2020_q2.lean | 4 +++ archive/imo/imo2021_q1.lean | 4 +++ .../2021summer/week3_p1.lean | 2 +- archive/sensitivity.lean | 4 +++ ...nically_ordered_comm_semiring_two_mul.lean | 4 +++ counterexamples/char_p_zero_ne_char_zero.lean | 4 +++ counterexamples/cyclotomic_105.lean | 4 +++ counterexamples/direct_sum_is_internal.lean | 4 +++ counterexamples/girard.lean | 4 +++ .../homogeneous_prime_not_prime.lean | 4 +++ ...linear_order_with_pos_mul_pos_eq_zero.lean | 4 +++ counterexamples/map_floor.lean | 4 +++ counterexamples/phillips.lean | 4 +++ counterexamples/pseudoelement.lean | 6 ++-- counterexamples/quadratic_form.lean | 4 +++ .../seminorm_lattice_not_distrib.lean | 4 +++ counterexamples/sorgenfrey_line.lean | 4 +++ .../zero_divisors_in_add_monoid_algebras.lean | 4 +++ 54 files changed, 226 insertions(+), 19 deletions(-) diff --git a/archive/100-theorems-list/37_solution_of_cubic.lean b/archive/100-theorems-list/37_solution_of_cubic.lean index aca8b61c940bc..8c5c11284e4e6 100644 --- a/archive/100-theorems-list/37_solution_of_cubic.lean +++ b/archive/100-theorems-list/37_solution_of_cubic.lean @@ -35,6 +35,8 @@ Originally ported from Isabelle/HOL. The polynomial, cubic, root -/ +namespace theorems_100 + section field open polynomial @@ -184,3 +186,5 @@ begin end end field + +end theorems_100 diff --git a/archive/100-theorems-list/42_inverse_triangle_sum.lean b/archive/100-theorems-list/42_inverse_triangle_sum.lean index 063c4d35bf0b7..0b12999fbf3d8 100644 --- a/archive/100-theorems-list/42_inverse_triangle_sum.lean +++ b/archive/100-theorems-list/42_inverse_triangle_sum.lean @@ -23,7 +23,7 @@ open_locale big_operators open finset /-- **Sum of the Reciprocals of the Triangular Numbers** -/ -lemma inverse_triangle_sum : +lemma theorem_100.inverse_triangle_sum : ∀ n, ∑ k in range n, (2 : ℚ) / (k * (k + 1)) = if n = 0 then 0 else 2 - (2 : ℚ) / n := begin refine sum_range_induction _ _ (if_pos rfl) _, diff --git a/archive/100-theorems-list/45_partition.lean b/archive/100-theorems-list/45_partition.lean index 0fe42d3f7c6f1..8796e3b3c2a4d 100644 --- a/archive/100-theorems-list/45_partition.lean +++ b/archive/100-theorems-list/45_partition.lean @@ -53,6 +53,8 @@ https://en.wikipedia.org/wiki/Partition_(number_theory)#Odd_parts_and_distinct_p -/ open power_series +namespace theorems_100 + noncomputable theory variables {α : Type*} @@ -521,3 +523,5 @@ begin rw odd_gf_prop n (n+1) (by linarith), apply same_coeffs (n+1) n n.le_succ, end + +end theorems_100 diff --git a/archive/100-theorems-list/57_herons_formula.lean b/archive/100-theorems-list/57_herons_formula.lean index 7d288df537ac0..100102469b88c 100644 --- a/archive/100-theorems-list/57_herons_formula.lean +++ b/archive/100-theorems-list/57_herons_formula.lean @@ -21,6 +21,8 @@ lengths. open real euclidean_geometry open_locale real euclidean_geometry +namespace theorems_100 + local notation `√` := real.sqrt variables {V : Type*} {P : Type*} @@ -61,3 +63,5 @@ begin one_mul, mul_div_cancel]; norm_num, end + +end theorems_100 diff --git a/archive/100-theorems-list/70_perfect_numbers.lean b/archive/100-theorems-list/70_perfect_numbers.lean index 69804193d4e8d..7d1cbd45655a7 100644 --- a/archive/100-theorems-list/70_perfect_numbers.lean +++ b/archive/100-theorems-list/70_perfect_numbers.lean @@ -26,21 +26,23 @@ Euler proved the converse, that if `n` is even and perfect, then there exists `k https://en.wikipedia.org/wiki/Euclid%E2%80%93Euler_theorem -/ +namespace theorems_100 + lemma odd_mersenne_succ (k : ℕ) : ¬ 2 ∣ mersenne (k + 1) := by simp [← even_iff_two_dvd, ← nat.even_add_one] with parity_simps namespace nat -open arithmetic_function finset +open nat.arithmetic_function finset open_locale arithmetic_function lemma sigma_two_pow_eq_mersenne_succ (k : ℕ) : σ 1 (2 ^ k) = mersenne (k + 1) := -by simp [sigma_one_apply, mersenne, prime_two, ← geom_sum_mul_add 1 (k+1)] +by simp [sigma_one_apply, mersenne, nat.prime_two, ← geom_sum_mul_add 1 (k+1)] /-- Euclid's theorem that Mersenne primes induce perfect numbers -/ theorem perfect_two_pow_mul_mersenne_of_prime (k : ℕ) (pr : (mersenne (k + 1)).prime) : - perfect ((2 ^ k) * mersenne (k + 1)) := + nat.perfect ((2 ^ k) * mersenne (k + 1)) := begin - rw [perfect_iff_sum_divisors_eq_two_mul, ← mul_assoc, ← pow_succ, ← sigma_one_apply, mul_comm, + rw [nat.perfect_iff_sum_divisors_eq_two_mul, ← mul_assoc, ← pow_succ, ← sigma_one_apply, mul_comm, is_multiplicative_sigma.map_mul_of_coprime (nat.prime_two.coprime_pow_of_not_dvd (odd_mersenne_succ _)), sigma_two_pow_eq_mersenne_succ], @@ -53,7 +55,7 @@ lemma ne_zero_of_prime_mersenne (k : ℕ) (pr : (mersenne (k + 1)).prime) : k ≠ 0 := begin intro H, - simpa [H, mersenne, not_prime_one] using pr, + simpa [H, mersenne, nat.not_prime_one] using pr, end theorem even_two_pow_mul_mersenne_of_prime (k : ℕ) (pr : (mersenne (k + 1)).prime) : @@ -77,26 +79,26 @@ end /-- **Perfect Number Theorem**: Euler's theorem that even perfect numbers can be factored as a power of two times a Mersenne prime. -/ -theorem eq_two_pow_mul_prime_mersenne_of_even_perfect {n : ℕ} (ev : even n) (perf : perfect n) : - ∃ (k : ℕ), prime (mersenne (k + 1)) ∧ n = 2 ^ k * mersenne (k + 1) := +theorem eq_two_pow_mul_prime_mersenne_of_even_perfect {n : ℕ} (ev : even n) (perf : nat.perfect n) : + ∃ (k : ℕ), nat.prime (mersenne (k + 1)) ∧ n = 2 ^ k * mersenne (k + 1) := begin have hpos := perf.2, rcases eq_two_pow_mul_odd hpos with ⟨k, m, rfl, hm⟩, use k, rw even_iff_two_dvd at hm, - rw [perfect_iff_sum_divisors_eq_two_mul hpos, ← sigma_one_apply, + rw [nat.perfect_iff_sum_divisors_eq_two_mul hpos, ← sigma_one_apply, is_multiplicative_sigma.map_mul_of_coprime (nat.prime_two.coprime_pow_of_not_dvd hm).symm, sigma_two_pow_eq_mersenne_succ, ← mul_assoc, ← pow_succ] at perf, rcases nat.coprime.dvd_of_dvd_mul_left (nat.prime_two.coprime_pow_of_not_dvd (odd_mersenne_succ _)) (dvd.intro _ perf) with ⟨j, rfl⟩, rw [← mul_assoc, mul_comm _ (mersenne _), mul_assoc] at perf, have h := mul_left_cancel₀ (ne_of_gt (mersenne_pos (nat.succ_pos _))) perf, - rw [sigma_one_apply, sum_divisors_eq_sum_proper_divisors_add_self, ← succ_mersenne, add_mul, + rw [sigma_one_apply, nat.sum_divisors_eq_sum_proper_divisors_add_self, ← succ_mersenne, add_mul, one_mul, add_comm] at h, have hj := add_left_cancel h, - cases sum_proper_divisors_dvd (by { rw hj, apply dvd.intro_left (mersenne (k + 1)) rfl }), + cases nat.sum_proper_divisors_dvd (by { rw hj, apply dvd.intro_left (mersenne (k + 1)) rfl }), { have j1 : j = 1 := eq.trans hj.symm h_1, - rw [j1, mul_one, sum_proper_divisors_eq_one_iff_prime] at h_1, + rw [j1, mul_one, nat.sum_proper_divisors_eq_one_iff_prime] at h_1, simp [h_1, j1] }, { have jcon := eq.trans hj.symm h_1, rw [← one_mul j, ← mul_assoc, mul_one] at jcon, @@ -108,7 +110,7 @@ begin rw [← jcon2, one_mul], exact even_iff_two_dvd.mp ev }, apply ne_of_lt _ jcon2, - rw [mersenne, ← nat.pred_eq_sub_one, lt_pred_iff, ← pow_one (nat.succ 1)], + rw [mersenne, ← nat.pred_eq_sub_one, nat.lt_pred_iff, ← pow_one (nat.succ 1)], apply pow_lt_pow (nat.lt_succ_self 1) (nat.succ_lt_succ (nat.succ_pos k)) }, contrapose! hm, simp [hm] } @@ -116,13 +118,16 @@ end /-- The Euclid-Euler theorem characterizing even perfect numbers -/ theorem even_and_perfect_iff {n : ℕ} : - (even n ∧ perfect n) ↔ ∃ (k : ℕ), prime (mersenne (k + 1)) ∧ n = 2 ^ k * mersenne (k + 1) := + (even n ∧ nat.perfect n) ↔ + ∃ (k : ℕ), nat.prime (mersenne (k + 1)) ∧ n = 2 ^ k * mersenne (k + 1) := begin split, { rintro ⟨ev, perf⟩, - exact eq_two_pow_mul_prime_mersenne_of_even_perfect ev perf }, + exact nat.eq_two_pow_mul_prime_mersenne_of_even_perfect ev perf }, { rintro ⟨k, pr, rfl⟩, exact ⟨even_two_pow_mul_mersenne_of_prime k pr, perfect_two_pow_mul_mersenne_of_prime k pr⟩ } end end nat + +end theorems_100 diff --git a/archive/100-theorems-list/73_ascending_descending_sequences.lean b/archive/100-theorems-list/73_ascending_descending_sequences.lean index ce47545e5e0c2..2c2d3931d8d69 100644 --- a/archive/100-theorems-list/73_ascending_descending_sequences.lean +++ b/archive/100-theorems-list/73_ascending_descending_sequences.lean @@ -26,6 +26,8 @@ variables {α : Type*} [linear_order α] {β : Type*} open function finset open_locale classical +namespace theorems_100 + /-- **Erdős–Szekeres Theorem**: Given a sequence of more than `r * s` distinct values, there is an increasing sequence of length longer than `r` or a decreasing sequence of length longer than `s`. @@ -158,3 +160,5 @@ begin -- Which follows from considering the cardinalities of the subset above, since `ab` is injective. simpa [nat.succ_injective, card_image_of_injective, ‹injective ab›] using card_le_of_subset this, end + +end theorems_100 diff --git a/archive/100-theorems-list/81_sum_of_prime_reciprocals_diverges.lean b/archive/100-theorems-list/81_sum_of_prime_reciprocals_diverges.lean index a07a1d585cafc..062c3d870d2d6 100644 --- a/archive/100-theorems-list/81_sum_of_prime_reciprocals_diverges.lean +++ b/archive/100-theorems-list/81_sum_of_prime_reciprocals_diverges.lean @@ -42,6 +42,8 @@ open_locale big_operators open_locale classical open filter finset +namespace theorems_100 + /-- The primes in `(k, x]`. -/ @@ -246,3 +248,5 @@ begin ... < x / 2 + x / 2 : add_lt_add_of_lt_of_le h3 h4 ... = x : add_halves ↑x, end + +end theorems_100 diff --git a/archive/100-theorems-list/82_cubing_a_cube.lean b/archive/100-theorems-list/82_cubing_a_cube.lean index 62061c0d53a07..ab4b567df3ac4 100644 --- a/archive/100-theorems-list/82_cubing_a_cube.lean +++ b/archive/100-theorems-list/82_cubing_a_cube.lean @@ -18,6 +18,8 @@ http://www.alaricstephen.com/main-featured/2017/9/28/cubing-a-cube-proof open real set function fin +namespace theorems_100 + noncomputable theory namespace «82» @@ -520,3 +522,5 @@ begin end end «82» + +end theorems_100 diff --git a/archive/100-theorems-list/83_friendship_graphs.lean b/archive/100-theorems-list/83_friendship_graphs.lean index a60e1e69bdf6a..adf5f345655a9 100644 --- a/archive/100-theorems-list/83_friendship_graphs.lean +++ b/archive/100-theorems-list/83_friendship_graphs.lean @@ -39,6 +39,8 @@ be phrased in terms of counting walks. -/ open_locale classical big_operators +namespace theorems_100 + noncomputable theory open finset simple_graph matrix @@ -144,7 +146,7 @@ begin use G.degree v, intro x, by_cases hvx : G.adj v x, swap, { exact (degree_eq_of_not_adj hG hvx).symm, }, - dunfold exists_politician at hG', + dunfold theorems_100.exists_politician at hG', push_neg at hG', rcases hG' v with ⟨w, hvw', hvw⟩, rcases hG' x with ⟨y, hxy', hxy⟩, @@ -336,3 +338,5 @@ begin { exact npG (hG.exists_politician_of_degree_le_two dreg (nat.lt_succ_iff.mp dle2)) }, { exact hG.false_of_three_le_degree dreg dge3 }, end + +end theorems_100 diff --git a/archive/100-theorems-list/93_birthday_problem.lean b/archive/100-theorems-list/93_birthday_problem.lean index 6d9c2d8bf3b1d..d4544a61f7f70 100644 --- a/archive/100-theorems-list/93_birthday_problem.lean +++ b/archive/100-theorems-list/93_birthday_problem.lean @@ -17,6 +17,8 @@ in terms of injective functions. The general result about `fintype.card (α ↪ uses is `fintype.card_embedding_eq`. -/ +namespace theorems_100 + local notation (name := finset.card) `|` x `|` := finset.card x local notation (name := fintype.card) `‖` x `‖` := fintype.card x @@ -75,3 +77,5 @@ begin end end measure_theory + +end theorems_100 diff --git a/archive/100-theorems-list/9_area_of_a_circle.lean b/archive/100-theorems-list/9_area_of_a_circle.lean index 76a34c96e6023..e3c8d7a0daf6f 100644 --- a/archive/100-theorems-list/9_area_of_a_circle.lean +++ b/archive/100-theorems-list/9_area_of_a_circle.lean @@ -43,6 +43,8 @@ to the n-ball. open set real measure_theory interval_integral open_locale real nnreal +namespace theorems_100 + /-- A disc of radius `r` is defined as the collection of points `(p.1, p.2)` in `ℝ × ℝ` such that `p.1 ^ 2 + p.2 ^ 2 < r ^ 2`. Note that this definition is not equivalent to `metric.ball (0 : ℝ × ℝ) r`. This was done @@ -115,3 +117,5 @@ begin hcont hderiv (continuous_const.mul hf).continuous_on.interval_integrable ... = nnreal.pi * r ^ 2 : by norm_num [F, inv_mul_cancel hlt.ne', ← mul_div_assoc, mul_comm π], end + +end theorems_100 diff --git a/archive/examples/prop_encodable.lean b/archive/examples/prop_encodable.lean index 4890ca2837d89..a877eeed41334 100644 --- a/archive/examples/prop_encodable.lean +++ b/archive/examples/prop_encodable.lean @@ -24,6 +24,8 @@ We mark the auxiliary constructions `private`, since their only purpose is to show encodability. -/ +namespace prop_encodable + /-- Propositional formulas with labels from `α`. -/ inductive prop_form (α : Type*) | var : α → prop_form @@ -96,3 +98,5 @@ begin end end prop_form + +end prop_encodable diff --git a/archive/imo/imo1959_q1.lean b/archive/imo/imo1959_q1.lean index 750f25078a713..b1af08fd9b9fa 100644 --- a/archive/imo/imo1959_q1.lean +++ b/archive/imo/imo1959_q1.lean @@ -18,6 +18,8 @@ as saying the numerator and denominator are relatively prime. open nat +namespace imo + lemma calculation (n k : ℕ) (h1 : k ∣ 21 * n + 4) (h2 : k ∣ 14 * n + 3) : k ∣ 1 := have h3 : k ∣ 2 * (21 * n + 4), from h1.mul_left 2, have h4 : k ∣ 3 * (14 * n + 3), from h2.mul_left 3, @@ -26,3 +28,5 @@ have h5 : 3 * (14 * n + 3) = 2 * (21 * n + 4) + 1, by ring, theorem imo1959_q1 : ∀ n : ℕ, coprime (21 * n + 4) (14 * n + 3) := assume n, coprime_of_dvd' $ λ k hp h1 h2, calculation n k h1 h2 + +end imo diff --git a/archive/imo/imo1962_q4.lean b/archive/imo/imo1962_q4.lean index 42bf061c48fb1..65b534aba58e9 100644 --- a/archive/imo/imo1962_q4.lean +++ b/archive/imo/imo1962_q4.lean @@ -17,6 +17,8 @@ in fact the simplest form of the set of solutions, and then prove it equals the open real open_locale real +namespace imo + noncomputable theory def problem_equation (x : ℝ) : Prop := cos x ^ 2 + cos (2 * x) ^ 2 + cos (3 * x) ^ 2 = 1 @@ -129,3 +131,5 @@ calc problem_equation x ↔ cos x ^ 2 + cos (2 * x) ^ 2 + cos (3 * x) ^ 2 = 1 : by refl ... ↔ cos (2 * x) = 0 ∨ cos (3 * x) = 0 : by simp [cos_two_mul, cos_three_mul, formula] ... ↔ x ∈ solution_set : by { rw [solve_cos2x_0, solve_cos3x_0, ← exists_or_distrib], refl } + +end imo diff --git a/archive/imo/imo1972_q5.lean b/archive/imo/imo1972_q5.lean index 21178464e58dd..27b0e32c8e3f0 100644 --- a/archive/imo/imo1972_q5.lean +++ b/archive/imo/imo1972_q5.lean @@ -15,6 +15,8 @@ Problem: `f` and `g` are real-valued functions defined on the real line. For all Prove that `|g(x)| ≤ 1` for all `x`. -/ +namespace imo + /-- This proof begins by introducing the supremum of `f`, `k ≤ 1` as well as `k' = k / ‖g y‖`. We then suppose that the conclusion does not hold (`hneg`) and show that `k ≤ k'` (by @@ -120,3 +122,5 @@ begin ... ≤ 2 * k : by linarith [h (x+y), h (x -y)] }, linarith, end + +end imo diff --git a/archive/imo/imo1975_q1.lean b/archive/imo/imo1975_q1.lean index 47170ea79e550..43a6e9409a65b 100644 --- a/archive/imo/imo1975_q1.lean +++ b/archive/imo/imo1975_q1.lean @@ -31,6 +31,8 @@ variables (hx : antitone_on x (finset.Icc 1 n)) variables (hy : antitone_on y (finset.Icc 1 n)) include hx hy hσ +namespace imo + theorem IMO_1975_Q1 : ∑ i in finset.Icc 1 n, (x i - y i) ^ 2 ≤ ∑ i in finset.Icc 1 n, (x i - y (σ i)) ^ 2 := begin @@ -46,3 +48,5 @@ begin -- them being `decreasing` exact antitone_on.monovary_on hx hy end + +end imo diff --git a/archive/imo/imo1977_q6.lean b/archive/imo/imo1977_q6.lean index 1c65f6af0c855..79342518718c7 100644 --- a/archive/imo/imo1977_q6.lean +++ b/archive/imo/imo1977_q6.lean @@ -15,6 +15,8 @@ We first prove the problem statement for `f : ℕ → ℕ` then we use it to prove the statement for positive naturals. -/ +namespace imo + theorem imo1977_q6_nat (f : ℕ → ℕ) (h : ∀ n, f (f n) < f (n + 1)) : ∀ n, f n = n := begin @@ -42,3 +44,5 @@ begin { simp }, { simpa using h _ } } end + +end imo diff --git a/archive/imo/imo1988_q6.lean b/archive/imo/imo1988_q6.lean index 044358ba8f7bd..8913a0c8dc30a 100644 --- a/archive/imo/imo1988_q6.lean +++ b/archive/imo/imo1988_q6.lean @@ -25,6 +25,8 @@ To illustrate the technique, we also prove a similar result. -- open_locale classical +namespace imo + local attribute [instance] classical.prop_decidable local attribute [simp] sq @@ -296,3 +298,5 @@ begin obtain rfl|rfl := (nat.dvd_prime nat.prime_two).mp y_dvd; apply mul_left_cancel₀, exacts [one_ne_zero, h.symm, two_ne_zero, h.symm] } } end + +end imo diff --git a/archive/imo/imo1994_q1.lean b/archive/imo/imo1994_q1.lean index bab8915e6af63..ffc8762b0e0d4 100644 --- a/archive/imo/imo1994_q1.lean +++ b/archive/imo/imo1994_q1.lean @@ -32,6 +32,8 @@ open_locale big_operators open finset +namespace imo + lemma tedious (m : ℕ) (k : fin (m+1)) : m - (m + (m + 1 - ↑k)) % (m + 1) = ↑k := begin cases k with k hk, @@ -97,3 +99,5 @@ begin -- A set of size `k+1` embed in one of size `k`, which yields a contradiction simpa [fin.coe_sub, tedious] using card_le_of_subset hf, end + +end imo diff --git a/archive/imo/imo1998_q2.lean b/archive/imo/imo1998_q2.lean index c4133b49d376e..928d78975e0d5 100644 --- a/archive/imo/imo1998_q2.lean +++ b/archive/imo/imo1998_q2.lean @@ -38,6 +38,8 @@ Rearranging gives the result. -/ open_locale classical +namespace imo + noncomputable theory /-- An ordered pair of judges. -/ @@ -212,3 +214,5 @@ begin { simp, }, { exact le_of_mul_le_mul_right h z.succ_pos, }, end + +end imo diff --git a/archive/imo/imo2001_q2.lean b/archive/imo/imo2001_q2.lean index a47afd9d539ac..bce2a35ab9b3e 100644 --- a/archive/imo/imo2001_q2.lean +++ b/archive/imo/imo2001_q2.lean @@ -30,6 +30,8 @@ open real variables {a b c : ℝ} +namespace imo + lemma denom_pos (ha : 0 < a) (hb : 0 < b) (hc : 0 < c) : 0 < a ^ 4 + b ^ 4 + c ^ 4 := add_pos (add_pos (pow_pos ha 4) (pow_pos hb 4)) (pow_pos hc 4) @@ -72,3 +74,5 @@ have h3 : ∀ {x : ℝ}, 0 < x → (x ^ (3 : ℝ)⁻¹) ^ 3 = x := λ x hx, show ↑3 = (3 : ℝ), by norm_num ▸ rpow_nat_inv_pow_nat hx.le three_ne_zero, calc 1 ≤ _ : imo2001_q2' (rpow_pos_of_pos ha _) (rpow_pos_of_pos hb _) (rpow_pos_of_pos hc _) ... = _ : by rw [h3 ha, h3 hb, h3 hc] + +end imo diff --git a/archive/imo/imo2001_q6.lean b/archive/imo/imo2001_q6.lean index c2fa9f829a855..05f09cbd59265 100644 --- a/archive/imo/imo2001_q6.lean +++ b/archive/imo/imo2001_q6.lean @@ -19,6 +19,8 @@ Prove that $a*b + c*d$ is not prime. variables {a b c d : ℤ} +namespace imo + theorem imo2001_q6 (hd : 0 < d) (hdc : d < c) (hcb : c < b) (hba : b < a) (h : a*c + b*d = (a + b - c + d) * (-a + b + c + d)) : ¬ prime (a*b + c*d) := @@ -42,3 +44,5 @@ begin have : a*c + b*d ≤ a*d + b*c, { from int.le_of_dvd aux h2 }, nlinarith only [hba, hdc, h, this] }, end + +end imo diff --git a/archive/imo/imo2005_q3.lean b/archive/imo/imo2005_q3.lean index 2c4acffbb7a11..b963c26b3ac78 100644 --- a/archive/imo/imo2005_q3.lean +++ b/archive/imo/imo2005_q3.lean @@ -18,6 +18,8 @@ factoring `(x^5-x^2)/(x^5+y^2+z^2) - (x^5-x^2)/(x^3*(x^2+y^2+z^2))` into a non-n and then making use of `xyz ≥ 1` to show `(x^5-x^2)/(x^3*(x^2+y^2+z^2)) ≥ (x^2-y*z)/(x^2+y^2+z^2)`. -/ +namespace imo + lemma key_insight (x y z : ℝ) (hx : x > 0) (hy : y > 0) (hz : z > 0) (h : x*y*z ≥ 1) : (x^5-x^2)/(x^5+y^2+z^2) ≥ (x^2-y*z)/(x^2+y^2+z^2) := begin @@ -59,3 +61,5 @@ begin (by linarith [sq_nonneg (x-y), sq_nonneg (y-z), sq_nonneg (z-x)]) (by linarith [sq_nonneg x, sq_nonneg y, sq_nonneg z]) }, end + +end imo diff --git a/archive/imo/imo2005_q4.lean b/archive/imo/imo2005_q4.lean index c6d854167e288..c31d8aeae822e 100644 --- a/archive/imo/imo2005_q4.lean +++ b/archive/imo/imo2005_q4.lean @@ -15,6 +15,8 @@ This is quite an easy problem, in which the key point is a modular arithmetic ca the sequence `a n` relative to an arbitrary prime. -/ +namespace imo + /-- The sequence considered in the problem, `2 ^ n + 3 ^ n + 6 ^ n - 1`. -/ def a (n : ℕ) : ℤ := 2 ^ n + 3 ^ n + 6 ^ n - 1 @@ -93,3 +95,5 @@ begin -- Contradiction! contradiction, end + +end imo diff --git a/archive/imo/imo2006_q3.lean b/archive/imo/imo2006_q3.lean index c6506dbcda082..9a266e1fb6a78 100644 --- a/archive/imo/imo2006_q3.lean +++ b/archive/imo/imo2006_q3.lean @@ -27,6 +27,8 @@ It involves making the substitution `x = a - b`, `y = b - c`, `z = c - a`, `s = a + b + c`. -/ +namespace imo + open real /-- Replacing `x` and `y` with their average increases the left side. -/ @@ -139,3 +141,5 @@ theorem imo2006_q3 (M : ℝ) : M * (a^2 + b^2 + c^2)^2) ↔ 9 * sqrt 2 / 32 ≤ M := ⟨proof₂ M, λ h _ _ _, le_trans proof₁ $ mul_le_mul_of_nonneg_right h $ sq_nonneg _⟩ + +end imo diff --git a/archive/imo/imo2006_q5.lean b/archive/imo/imo2006_q5.lean index 665882fd84f23..e170e157a1798 100644 --- a/archive/imo/imo2006_q5.lean +++ b/archive/imo/imo2006_q5.lean @@ -41,8 +41,10 @@ $P(t)+t-a-b$, and we're again done. open function polynomial +namespace imo /-- If every entry in a cyclic list of integers divides the next, then they all have the same absolute value. -/ + theorem int.nat_abs_eq_of_chain_dvd {l : cycle ℤ} {x y : ℤ} (hl : l.chain (∣)) (hx : x ∈ l) (hy : y ∈ l) : x.nat_abs = y.nat_abs := begin @@ -208,3 +210,5 @@ begin simp only [sub_eq_zero, is_root.def, eval_sub, iterate_comp_eval, eval_X] at ht, simpa [mem_roots hP', sub_eq_zero] using polynomial.is_periodic_pt_eval_two ⟨k, hk, ht⟩ end + +end imo diff --git a/archive/imo/imo2008_q2.lean b/archive/imo/imo2008_q2.lean index a20f897827eef..f633e57bbf8a9 100644 --- a/archive/imo/imo2008_q2.lean +++ b/archive/imo/imo2008_q2.lean @@ -26,6 +26,8 @@ using `c`, `m` and `n`. We factor `LHS - 1` as a square, which finishes the proo set of rational solutions to the equation, and that `W` is infinite. -/ +namespace imo + lemma subst_abc {x y z : ℝ} (h : x*y*z = 1) : ∃ a b c : ℝ, a ≠ 0 ∧ b ≠ 0 ∧ c ≠ 0 ∧ x = a/b ∧ y = b/c ∧ z = c /a := begin @@ -132,3 +134,5 @@ begin exact hW_inf.mono hW_sub_S, end + +end imo diff --git a/archive/imo/imo2008_q3.lean b/archive/imo/imo2008_q3.lean index 3f8d05c99f40f..a86b5ab4fe95a 100644 --- a/archive/imo/imo2008_q3.lean +++ b/archive/imo/imo2008_q3.lean @@ -28,6 +28,8 @@ Then `p = 2n + k ≥ 2n + √(p - 4) = 2n + √(2n + k - 4) > √(2n)` and we ar open real +namespace imo + lemma p_lemma (p : ℕ) (hpp : nat.prime p) (hp_mod_4_eq_1 : p ≡ 1 [MOD 4]) (hp_gt_20 : p > 20) : ∃ n : ℕ, p ∣ n ^ 2 + 1 ∧ (p : ℝ) > 2 * n + sqrt(2 * n) := begin @@ -90,3 +92,5 @@ begin exact ⟨n, hn_ge_N, p, hpp, hnat, hreal⟩, end + +end imo diff --git a/archive/imo/imo2008_q4.lean b/archive/imo/imo2008_q4.lean index b82ea5691640d..cc9889a6dec1a 100644 --- a/archive/imo/imo2008_q4.lean +++ b/archive/imo/imo2008_q4.lean @@ -23,6 +23,8 @@ The desired theorem is that either `f = λ x, x` or `f = λ x, 1/x` open real +namespace imo + lemma abs_eq_one_of_pow_eq_one (x : ℝ) (n : ℕ) (hn : n ≠ 0) (h : x ^ n = 1) : |x| = 1 := by rw [← pow_left_inj (abs_nonneg x) zero_le_one (pos_iff_ne_zero.2 hn), one_pow, pow_abs, h, abs_one] @@ -103,3 +105,5 @@ begin obtain ha₂ := abs_eq_one_of_pow_eq_one a 4 (show 4 ≠ 0, by norm_num) ha₁, rw abs_of_pos ha at ha₂, rw ha₂ at hfa₁, norm_num at hfa₁ }, end + +end imo diff --git a/archive/imo/imo2011_q3.lean b/archive/imo/imo2011_q3.lean index c3e5dfc5d8361..1889e503ef54b 100644 --- a/archive/imo/imo2011_q3.lean +++ b/archive/imo/imo2011_q3.lean @@ -20,6 +20,8 @@ for all x and y. Prove that f(x) = 0 for all x ≤ 0. Direct translation of the solution found in https://www.imo-official.org/problems/IMO2011SL.pdf -/ +namespace imo + theorem imo2011_q3 (f : ℝ → ℝ) (hf : ∀ x y, f (x + y) ≤ y * f x + f (f x)) : ∀ x ≤ 0, f x = 0 := begin @@ -68,3 +70,5 @@ begin rw hno at hp, linarith }, end + +end imo diff --git a/archive/imo/imo2011_q5.lean b/archive/imo/imo2011_q5.lean index abf6897ad352d..5fe984f6114a6 100644 --- a/archive/imo/imo2011_q5.lean +++ b/archive/imo/imo2011_q5.lean @@ -17,6 +17,8 @@ of positive integers. Suppose that, for any two integers open int +namespace imo + theorem imo2011_q5 (f : ℤ → ℤ) (hpos : ∀ n : ℤ, 0 < f n) (hdvd : ∀ m n : ℤ, f (m - n) ∣ f m - f n) : ∀ m n : ℤ, f m ≤ f n → f m ∣ f n := @@ -57,3 +59,5 @@ begin { -- m = n rw h_fm_eq_fn } end + +end imo diff --git a/archive/imo/imo2013_q1.lean b/archive/imo/imo2013_q1.lean index 0b6b4680cc777..a2b0315d2f37e 100644 --- a/archive/imo/imo2013_q1.lean +++ b/archive/imo/imo2013_q1.lean @@ -27,6 +27,8 @@ We prove a slightly more general version where k does not need to be strictly po open_locale big_operators +namespace imo + lemma arith_lemma (k n : ℕ) : 0 < 2 * n + 2^k.succ := calc 0 < 2 : zero_lt_two ... = 2^1 : (pow_one 2).symm @@ -101,3 +103,5 @@ begin (1 + 1 / ↑(m pk)) : by rw [prod_lemma, hpm, ←hmpk, mul_comm] ... = ∏ i in finset.range pk.succ, (1 + 1 / m i) : by rw ← finset.prod_range_succ _ pk } end + +end imo diff --git a/archive/imo/imo2013_q5.lean b/archive/imo/imo2013_q5.lean index 53cb83b0df515..d27aa509953e8 100644 --- a/archive/imo/imo2013_q5.lean +++ b/archive/imo/imo2013_q5.lean @@ -29,6 +29,8 @@ https://www.imo-official.org/problems/IMO2013SL.pdf open_locale big_operators +namespace imo + lemma le_of_all_pow_lt_succ {x y : ℝ} (hx : 1 < x) (hy : 1 < y) (h : ∀ n : ℕ, 0 < n → x^n - 1 < y^n) : x ≤ y := @@ -301,3 +303,5 @@ begin ... = (((x2num : ℚ) / (x2denom : ℚ) : ℚ) : ℝ) : by norm_cast ... = x : by rw ←hrat_expand2 end + +end imo diff --git a/archive/imo/imo2019_q1.lean b/archive/imo/imo2019_q1.lean index b19fd9475f693..46d3bffacab40 100644 --- a/archive/imo/imo2019_q1.lean +++ b/archive/imo/imo2019_q1.lean @@ -19,6 +19,8 @@ Note that there is a much more compact proof of this fact in Isabelle/HOL - http://downthetypehole.de/paste/4YbGgqb4 -/ +namespace imo + theorem imo2019Q1 (f : ℤ → ℤ) : (∀ a b : ℤ, f (2 * a) + 2 * (f b) = f (f (a + b))) ↔ (f = 0) ∨ ∃ c, f = λ x, 2 * x + c := @@ -49,3 +51,5 @@ begin { right, use c, ext b, simp [H, add_comm] }, { left, ext b, simpa [H, two_ne_zero] using H3 } end + +end imo diff --git a/archive/imo/imo2019_q2.lean b/archive/imo/imo2019_q2.lean index ca6781c82a7e5..35d005480df5f 100644 --- a/archive/imo/imo2019_q2.lean +++ b/archive/imo/imo2019_q2.lean @@ -56,6 +56,8 @@ rather than more literally with `affine_segment`. -/ library_note "IMO geometry formalization conventions" +namespace imo + noncomputable theory open affine affine.simplex euclidean_geometry finite_dimensional @@ -591,3 +593,5 @@ theorem imo2019_q2 (A B C A₁ B₁ P Q P₁ Q₁ : Pt) (⟨A, B, C, A₁, B₁, P, Q, P₁, Q₁, affine_independent_ABC, wbtw_B_A₁_C, wbtw_A_B₁_C, wbtw_A_P_A₁, wbtw_B_Q_B₁, PQ_parallel_AB, P_ne_Q, sbtw_P_B₁_P₁, angle_PP₁C_eq_angle_BAC, C_ne_P₁, sbtw_Q_A₁_Q₁, angle_CQ₁Q_eq_angle_CBA, C_ne_Q₁⟩ : imo2019q2_cfg V Pt).result + +end imo diff --git a/archive/imo/imo2019_q4.lean b/archive/imo/imo2019_q4.lean index a4d25853dab7f..dfc00ea7b44d8 100644 --- a/archive/imo/imo2019_q4.lean +++ b/archive/imo/imo2019_q4.lean @@ -26,6 +26,8 @@ individually. open_locale nat big_operators open finset multiplicity nat (hiding zero_le prime) +namespace imo + theorem imo2019_q4_upper_bound {k n : ℕ} (hk : k > 0) (h : (k! : ℤ) = ∏ i in range n, (2 ^ n - 2 ^ i)) : n < 6 := begin @@ -92,3 +94,5 @@ begin /- n = 5 -/ { refine monotone_factorial.ne_of_lt_of_lt_nat 10 _ _ _ h; norm_num }, end + +end imo diff --git a/archive/imo/imo2020_q2.lean b/archive/imo/imo2020_q2.lean index beea58a78ff17..4bc2f471ae33d 100644 --- a/archive/imo/imo2020_q2.lean +++ b/archive/imo/imo2020_q2.lean @@ -21,6 +21,8 @@ the official solutions. open real +namespace imo + theorem imo2020_q2 (a b c d : ℝ) (hd0 : 0 < d) (hdc : d ≤ c) (hcb : c ≤ b) (hba : b ≤ a) (h1 : a + b + c + d = 1) : (a + 2 * b + 3 * c + 4 * d) * a ^ a * b ^ b * c ^ c * d ^ d < 1 := @@ -44,3 +46,5 @@ begin ... = (a + b + c + d) ^ 3 : by ring ... = 1 : by simp [h1] end + +end imo diff --git a/archive/imo/imo2021_q1.lean b/archive/imo/imo2021_q1.lean index 4a123aeaa0f51..4a8f457c42722 100644 --- a/archive/imo/imo2021_q1.lean +++ b/archive/imo/imo2021_q1.lean @@ -43,6 +43,8 @@ which finishes the proof. open real +namespace imo + lemma lower_bound (n l : ℕ) (hl : 2 + sqrt (4 + 2 * n) ≤ 2 * l) : n + 4 * l ≤ 2 * l * l := begin @@ -178,3 +180,5 @@ begin cases hCA; [right, left]; exact ⟨a, (hCA ha).2, b, (hCA hb).2, hab, h₁ a (hCA ha).1 b (hCA hb).1 hab⟩, end + +end imo diff --git a/archive/oxford_invariants/2021summer/week3_p1.lean b/archive/oxford_invariants/2021summer/week3_p1.lean index d0a8fb0731ff8..eb30161c3a81c 100644 --- a/archive/oxford_invariants/2021summer/week3_p1.lean +++ b/archive/oxford_invariants/2021summer/week3_p1.lean @@ -72,7 +72,7 @@ open_locale big_operators variables {α : Type*} [linear_ordered_field α] -theorem week3_p1 (n : ℕ) (a : ℕ → ℕ) (a_pos : ∀ i ≤ n, 0 < a i) +theorem oxford_invariants.week3_p1 (n : ℕ) (a : ℕ → ℕ) (a_pos : ∀ i ≤ n, 0 < a i) (ha : ∀ i, i + 2 ≤ n → a (i + 1) ∣ a i + a (i + 2)) : ∃ b : ℕ, (b : α) = ∑ i in finset.range n, (a 0 * a n)/(a i * a (i + 1)) := begin diff --git a/archive/sensitivity.lean b/archive/sensitivity.lean index c2c46c393fd79..995ff612ae9c3 100644 --- a/archive/sensitivity.lean +++ b/archive/sensitivity.lean @@ -33,6 +33,8 @@ The project was developed at https://github.com/leanprover-community/lean-sensit archived at https://github.com/leanprover-community/mathlib/blob/master/archive/sensitivity.lean -/ +namespace sensitivity + /-! The next two lines assert we do not want to give a constructive proof, but rather use classical logic. -/ noncomputable theory @@ -435,3 +437,5 @@ begin convert finset.inter_subset_inter_right coeffs_support end end + +end sensitivity diff --git a/counterexamples/canonically_ordered_comm_semiring_two_mul.lean b/counterexamples/canonically_ordered_comm_semiring_two_mul.lean index fd6ca5897584d..012d49af2dcee 100644 --- a/counterexamples/canonically_ordered_comm_semiring_two_mul.lean +++ b/counterexamples/canonically_ordered_comm_semiring_two_mul.lean @@ -23,6 +23,8 @@ Reference: https://leanprover.zulipchat.com/#narrow/stream/113489-new-members/topic/canonically_ordered.20pathology -/ +namespace counterexample + namespace from_Bhavik /-- Bhavik Mehta's example. There are only the initial definitions, but no proofs. The Type @@ -265,3 +267,5 @@ begin end end ex_L + +end counterexample diff --git a/counterexamples/char_p_zero_ne_char_zero.lean b/counterexamples/char_p_zero_ne_char_zero.lean index 36fd17d6adcf9..a6b9cecde95cc 100644 --- a/counterexamples/char_p_zero_ne_char_zero.lean +++ b/counterexamples/char_p_zero_ne_char_zero.lean @@ -18,6 +18,8 @@ This file shows that there are semiring `R` for which `char_p R 0` holds and `ch The example is `{0, 1}` with saturating addition. -/ +namespace counterexample + @[simp] lemma add_one_eq_one (x : with_zero unit) : x + 1 = 1 := with_zero.cases_on x (by refl) (λ h, by refl) @@ -26,3 +28,5 @@ lemma with_zero_unit_char_p_zero : char_p (with_zero unit) 0 := lemma with_zero_unit_not_char_zero : ¬ char_zero (with_zero unit) := λ ⟨h⟩, h.ne (by simp : 1 + 1 ≠ 0 + 1) (by simp) + +end counterexample diff --git a/counterexamples/cyclotomic_105.lean b/counterexamples/cyclotomic_105.lean index d600a7098a3d5..cfbb5aa23bb94 100644 --- a/counterexamples/cyclotomic_105.lean +++ b/counterexamples/cyclotomic_105.lean @@ -16,6 +16,8 @@ theorem `not_forall_coeff_cyclotomic_neg_one_zero_one`. We prove this with the c open nat (proper_divisors) finset +namespace counterexample + section computation instance nat.fact_prime_five : fact (nat.prime 5) := ⟨by norm_num⟩ @@ -100,3 +102,5 @@ begin rw coeff_cyclotomic_105 at h, norm_num at h end + +end counterexample diff --git a/counterexamples/direct_sum_is_internal.lean b/counterexamples/direct_sum_is_internal.lean index c9febdae9b63d..9a4b75ece1807 100644 --- a/counterexamples/direct_sum_is_internal.lean +++ b/counterexamples/direct_sum_is_internal.lean @@ -19,6 +19,8 @@ This file demonstrates why `direct_sum.is_internal_submodule_of_independent_of_s take `ring R` and not `semiring R`. -/ +namespace counterexample + lemma units_int.one_ne_neg_one : (1 : ℤˣ) ≠ -1 := dec_trivial /-- Submodules of positive and negative integers, keyed by sign. -/ @@ -93,3 +95,5 @@ end /-- And so they do not represent an internal direct sum. -/ lemma with_sign.not_internal : ¬direct_sum.is_internal with_sign := with_sign.not_injective ∘ and.elim_left + +end counterexample diff --git a/counterexamples/girard.lean b/counterexamples/girard.lean index 218c892caa71e..090551161b38a 100644 --- a/counterexamples/girard.lean +++ b/counterexamples/girard.lean @@ -24,6 +24,8 @@ Based on Watkins' LF implementation of Hurkens' simplification of Girard's parad * `girard`: there are no Girard universes. -/ +namespace counterexample + /-- **Girard's paradox**: there are no universes `u` such that `Type u : Type u`. Since we can't actually change the type of Lean's `Π` operator, we assume the existence of `pi`, `lam`, `app` and the `beta` rule equivalent to the `Π` and `app` constructors of type theory. @@ -43,3 +45,5 @@ let ω : set (set U) := {p | ∀ x, p ∈ σ x → x ∈ p} in let δ (S : set (set U)) := ∀ p, p ∈ S → τ S ∈ p in have δ ω := λ p d, d (τ ω) $ στ.2 $ λ x h, d (τ (σ x)) (στ.2 h), this {y | ¬ δ (σ y)} (λ x e f, f _ e (λ p h, f _ (στ.1 h))) (λ p h, this _ (στ.1 h)) + +end counterexample diff --git a/counterexamples/homogeneous_prime_not_prime.lean b/counterexamples/homogeneous_prime_not_prime.lean index 8692774542313..1e895db644e53 100644 --- a/counterexamples/homogeneous_prime_not_prime.lean +++ b/counterexamples/homogeneous_prime_not_prime.lean @@ -30,6 +30,8 @@ prime. But it is homogeneously prime, i.e. if `(a, b), (c, d)` are two homogeneo homogeneous, prime -/ +namespace counterexample + namespace counterexample_not_prime_but_homogeneous_prime open direct_sum @@ -152,3 +154,5 @@ begin end end counterexample_not_prime_but_homogeneous_prime + +end counterexample diff --git a/counterexamples/linear_order_with_pos_mul_pos_eq_zero.lean b/counterexamples/linear_order_with_pos_mul_pos_eq_zero.lean index 6e08c5f9f847f..0fac502d4e5e3 100644 --- a/counterexamples/linear_order_with_pos_mul_pos_eq_zero.lean +++ b/counterexamples/linear_order_with_pos_mul_pos_eq_zero.lean @@ -18,6 +18,8 @@ Relevant Zulip chat: https://leanprover.zulipchat.com/#narrow/stream/116395-maths/topic/mul_pos -/ +namespace counterexample + /-- The three element monoid. -/ @[derive [decidable_eq]] inductive foo @@ -85,3 +87,5 @@ end example : 0 < ε ∧ ε * ε = 0 := by boom end foo + +end counterexample diff --git a/counterexamples/map_floor.lean b/counterexamples/map_floor.lean index b1c262d6d780c..2fd1b31c7cedc 100644 --- a/counterexamples/map_floor.lean +++ b/counterexamples/map_floor.lean @@ -33,6 +33,8 @@ But it does not preserve floors (nor ceils) as `⌊-ε⌋ = -1` while `⌊f (-ε (`int_with_epsilons.forget_epsilons_floor_lt`, `int_with_epsilons.lt_forget_epsilons_ceil`). -/ +namespace counterexample + noncomputable theory open function int polynomial @@ -124,3 +126,5 @@ begin end end int_with_epsilon + +end counterexample diff --git a/counterexamples/phillips.lean b/counterexamples/phillips.lean index 82bdaa79e5d5d..eafafd6d169bd 100644 --- a/counterexamples/phillips.lean +++ b/counterexamples/phillips.lean @@ -72,6 +72,8 @@ on a discrete copy of the original type, as mathlib only contains the space of a continuous functions (which is the useful one). -/ +namespace counterexample + universe u variables {α : Type u} open set bounded_continuous_function measure_theory @@ -606,3 +608,5 @@ begin end end phillips_1940 + +end counterexample diff --git a/counterexamples/pseudoelement.lean b/counterexamples/pseudoelement.lean index 5a28cf838eabb..0e2ddaf8cc23b 100644 --- a/counterexamples/pseudoelement.lean +++ b/counterexamples/pseudoelement.lean @@ -31,9 +31,11 @@ given by `t ↦ (t, 2 * t)` and `y : ℚ ⟶ ℚ ⊞ ℚ` given by `t ↦ (t, t) open category_theory.abelian category_theory category_theory.limits Module linear_map +namespace counterexample + noncomputable theory -namespace category_theory.abelian.pseudoelement +open category_theory.abelian.pseudoelement /-- `x` is given by `t ↦ (t, 2 * t)`. -/ def x : over ((of ℤ ℚ) ⊞ (of ℤ ℚ)) := @@ -119,4 +121,4 @@ lemma exist_ne_and_fst_eq_fst_and_snd_eq_snd : ∃ x y : (of ℤ ℚ) ⊞ (of (biprod.snd : (of ℤ ℚ) ⊞ (of ℤ ℚ) ⟶ _) x = (biprod.snd : (of ℤ ℚ) ⊞ (of ℤ ℚ) ⟶ _) y:= ⟨⟦x⟧, ⟦y⟧, mk_x_ne_mk_y, fst_mk_x_eq_fst_mk_y, snd_mk_x_eq_snd_mk_y⟩ -end category_theory.abelian.pseudoelement +end counterexample diff --git a/counterexamples/quadratic_form.lean b/counterexamples/quadratic_form.lean index 20af196c19f52..d9f029822dce1 100644 --- a/counterexamples/quadratic_form.lean +++ b/counterexamples/quadratic_form.lean @@ -19,6 +19,8 @@ variables (F : Type*) [nontrivial F] [comm_ring F] [char_p F 2] open bilin_form +namespace counterexample + /-- The bilinear form we will use as a counterexample, over some field `F` of characteristic two. -/ def B : bilin_form F (F × F) := bilin_form.lin_mul_lin (linear_map.fst _ _ _) (linear_map.snd _ _ _) @@ -51,3 +53,5 @@ begin rw [bilin_form.to_quadratic_form_zero, bilin_form.to_quadratic_form_eq_zero], exact is_alt_B F, end + +end counterexample diff --git a/counterexamples/seminorm_lattice_not_distrib.lean b/counterexamples/seminorm_lattice_not_distrib.lean index 7963e3e7a07ef..95b4e997ea397 100644 --- a/counterexamples/seminorm_lattice_not_distrib.lean +++ b/counterexamples/seminorm_lattice_not_distrib.lean @@ -20,6 +20,8 @@ This proves the lattice `seminorm ℝ (ℝ × ℝ)` is not distributive. open seminorm open_locale nnreal +namespace counterexample + namespace seminorm_not_distrib @[simps] noncomputable def p : seminorm ℝ (ℝ×ℝ) := @@ -68,3 +70,5 @@ begin end end seminorm_not_distrib + +end counterexample diff --git a/counterexamples/sorgenfrey_line.lean b/counterexamples/sorgenfrey_line.lean index f96879bbd74c9..065ecbecd57b5 100644 --- a/counterexamples/sorgenfrey_line.lean +++ b/counterexamples/sorgenfrey_line.lean @@ -32,6 +32,8 @@ Prove that the Sorgenfrey line is a paracompact space. open set filter topological_space open_locale topology filter +namespace counterexample + noncomputable theory /-- The Sorgenfrey line. It is the real line with the topology space structure generated by @@ -294,3 +296,5 @@ lemma not_second_countable_topology : ¬second_countable_topology ℝₗ := by { introI, exact not_metrizable_space (metrizable_space_of_t3_second_countable _) } end sorgenfrey_line + +end counterexample diff --git a/counterexamples/zero_divisors_in_add_monoid_algebras.lean b/counterexamples/zero_divisors_in_add_monoid_algebras.lean index da188dfa571cb..88e430e411b30 100644 --- a/counterexamples/zero_divisors_in_add_monoid_algebras.lean +++ b/counterexamples/zero_divisors_in_add_monoid_algebras.lean @@ -41,6 +41,8 @@ finitely supported function is lexicographic, matching the list notation. The i -/ open finsupp add_monoid_algebra +namespace counterexample + /-- This is a simple example showing that if `R` is a non-trivial ring and `A` is an additive monoid with an element `a` satisfying `n • a = a` and `(n - 1) • a ≠ a`, for some `2 ≤ n`, then `add_monoid_algebra R A` contains non-zero zero-divisors. The elements are easy to write down: @@ -230,3 +232,5 @@ begin { simpa [unique_add] }, exact λ x y, ⟨x - 1, y + 1, sub_add_add_cancel _ _ _, by simp⟩, end + +end counterexample From 12a85fac627bea918960da036049d611b1a3ee43 Mon Sep 17 00:00:00 2001 From: Riccardo Brasca Date: Wed, 31 May 2023 00:27:53 +0000 Subject: [PATCH 1118/1291] chore(field_theory/finite/basic): move a lemma (#19130) Moving one instance we can erase `field_theory.splitting_field` from the imports of `field_theory.finite.basic)`. --- src/field_theory/finite/basic.lean | 20 -------------------- src/field_theory/finite/galois_field.lean | 20 +++++++++++++++++++- 2 files changed, 19 insertions(+), 21 deletions(-) diff --git a/src/field_theory/finite/basic.lean b/src/field_theory/finite/basic.lean index 1df22603ad038..75b448e404c72 100644 --- a/src/field_theory/finite/basic.lean +++ b/src/field_theory/finite/basic.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes, Joey van Langen, Casper Putz -/ import field_theory.separable -import field_theory.splitting_field import ring_theory.integral_domain import tactic.apply_fun @@ -219,7 +218,6 @@ begin ... = 0 : by { rw [sum_pow_units K i, if_neg], exact hiq, } end -section is_splitting_field open polynomial section @@ -268,24 +266,6 @@ begin zero_sub] }, end -instance (F : Type*) [field F] [algebra F K] : is_splitting_field F K (X^q - X) := -{ splits := - begin - have h : (X^q - X : K[X]).nat_degree = q := - X_pow_card_sub_X_nat_degree_eq K fintype.one_lt_card, - rw [←splits_id_iff_splits, splits_iff_card_roots, polynomial.map_sub, polynomial.map_pow, - map_X, h, roots_X_pow_card_sub_X K, ←finset.card_def, finset.card_univ], - end, - adjoin_roots := - begin - classical, - transitivity algebra.adjoin F ((roots (X^q - X : K[X])).to_finset : set K), - { simp only [polynomial.map_pow, map_X, polynomial.map_sub], }, - { rw [roots_X_pow_card_sub_X, val_to_finset, coe_univ, algebra.adjoin_univ], } - end } - -end is_splitting_field - variables {K} theorem frobenius_pow {p : ℕ} [fact p.prime] [char_p K p] {n : ℕ} (hcard : q = p^n) : diff --git a/src/field_theory/finite/galois_field.lean b/src/field_theory/finite/galois_field.lean index ac9250227e85c..9165891e346a9 100644 --- a/src/field_theory/finite/galois_field.lean +++ b/src/field_theory/finite/galois_field.lean @@ -8,6 +8,7 @@ import algebra.char_p.algebra import data.zmod.algebra import field_theory.finite.basic import field_theory.galois +import field_theory.splitting_field /-! # Galois fields @@ -30,9 +31,26 @@ It is a finite field with `p ^ n` elements. noncomputable theory -open polynomial +open polynomial finset open_locale polynomial +instance finite_field.has_sub.sub.polynomial.is_splitting_field (K F : Type*) [field K] [fintype K] + [field F] [algebra F K] : is_splitting_field F K (X ^ (fintype.card K) - X) := +{ splits := + begin + have h : (X ^ (fintype.card K) - X : K[X]).nat_degree = fintype.card K := + finite_field.X_pow_card_sub_X_nat_degree_eq K fintype.one_lt_card, + rw [←splits_id_iff_splits, splits_iff_card_roots, polynomial.map_sub, polynomial.map_pow, + map_X, h, finite_field.roots_X_pow_card_sub_X K, ←finset.card_def, finset.card_univ], + end, + adjoin_roots := + begin + classical, + transitivity algebra.adjoin F ((roots (X ^ (fintype.card K) - X : K[X])).to_finset : set K), + { simp only [polynomial.map_pow, map_X, polynomial.map_sub], }, + { rw [finite_field.roots_X_pow_card_sub_X, val_to_finset, coe_univ, algebra.adjoin_univ], } + end } + lemma galois_poly_separable {K : Type*} [field K] (p q : ℕ) [char_p K p] (h : p ∣ q) : separable (X ^ q - X : K[X]) := begin From 61b5e2755ccb464b68d05a9acf891ae04992d09d Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Wed, 31 May 2023 08:02:01 +0000 Subject: [PATCH 1119/1291] chore(*): add mathlib4 synchronization comments (#19132) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.category.Module.adjunctions` * `analysis.calculus.darboux` * `analysis.calculus.deriv.inv` * `analysis.calculus.diff_cont_on_cl` * `analysis.calculus.dslope` * `analysis.inner_product_space.gram_schmidt_ortho` * `category_theory.abelian.ext` * `category_theory.abelian.left_derived` * `category_theory.bicategory.single_obj` * `data.real.golden_ratio` * `measure_theory.category.Meas` * `measure_theory.decomposition.jordan` * `measure_theory.decomposition.signed_hahn` * `measure_theory.group.action` * `measure_theory.group.measure` * `number_theory.arithmetic_function` * `number_theory.l_series` * `number_theory.von_mangoldt` * `topology.algebra.continuous_monoid_hom` * `topology.instances.discrete` * `topology.sheaves.stalks` * `topology.tietze_extension` --- src/algebra/category/Module/adjunctions.lean | 3 +++ src/analysis/calculus/darboux.lean | 3 +++ src/analysis/calculus/deriv/inv.lean | 3 +++ src/analysis/calculus/diff_cont_on_cl.lean | 3 +++ src/analysis/calculus/dslope.lean | 3 +++ src/analysis/inner_product_space/gram_schmidt_ortho.lean | 3 +++ src/category_theory/abelian/ext.lean | 3 +++ src/category_theory/abelian/left_derived.lean | 3 +++ src/category_theory/bicategory/single_obj.lean | 3 +++ src/data/real/golden_ratio.lean | 3 +++ src/measure_theory/category/Meas.lean | 3 +++ src/measure_theory/decomposition/jordan.lean | 3 +++ src/measure_theory/decomposition/signed_hahn.lean | 3 +++ src/measure_theory/group/action.lean | 3 +++ src/measure_theory/group/measure.lean | 3 +++ src/number_theory/arithmetic_function.lean | 3 +++ src/number_theory/l_series.lean | 3 +++ src/number_theory/von_mangoldt.lean | 3 +++ src/topology/algebra/continuous_monoid_hom.lean | 3 +++ src/topology/instances/discrete.lean | 3 +++ src/topology/sheaves/stalks.lean | 3 +++ src/topology/tietze_extension.lean | 3 +++ 22 files changed, 66 insertions(+) diff --git a/src/algebra/category/Module/adjunctions.lean b/src/algebra/category/Module/adjunctions.lean index b2d2ceb2793cf..0a2b07344bc18 100644 --- a/src/algebra/category/Module/adjunctions.lean +++ b/src/algebra/category/Module/adjunctions.lean @@ -10,6 +10,9 @@ import linear_algebra.direct_sum.finsupp import category_theory.linear.linear_functor /-! +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The functor of forming finitely supported functions on a type with values in a `[ring R]` is the left adjoint of the forgetful functor from `R`-modules to types. diff --git a/src/analysis/calculus/darboux.lean b/src/analysis/calculus/darboux.lean index 529cfb0bf3e91..988217525e49e 100644 --- a/src/analysis/calculus/darboux.lean +++ b/src/analysis/calculus/darboux.lean @@ -8,6 +8,9 @@ import analysis.calculus.local_extr /-! # Darboux's theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove that the derivative of a differentiable function on an interval takes all intermediate values. The proof is based on the [Wikipedia](https://en.wikipedia.org/wiki/Darboux%27s_theorem_(analysis)) page about this theorem. diff --git a/src/analysis/calculus/deriv/inv.lean b/src/analysis/calculus/deriv/inv.lean index ec37caff033d7..5895b4209ab95 100644 --- a/src/analysis/calculus/deriv/inv.lean +++ b/src/analysis/calculus/deriv/inv.lean @@ -9,6 +9,9 @@ import analysis.calculus.deriv.comp /-! # Derivatives of `x ↦ x⁻¹` and `f x / g x` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove `(x⁻¹)' = -1 / x ^ 2`, `((f x)⁻¹)' = -f' x / (f x) ^ 2`, and `(f x / g x)' = (f' x * g x - f x * g' x) / (g x) ^ 2` for different notions of derivative. diff --git a/src/analysis/calculus/diff_cont_on_cl.lean b/src/analysis/calculus/diff_cont_on_cl.lean index 4150d47e822da..7c9c611b29ebf 100644 --- a/src/analysis/calculus/diff_cont_on_cl.lean +++ b/src/analysis/calculus/diff_cont_on_cl.lean @@ -8,6 +8,9 @@ import analysis.calculus.deriv.inv /-! # Functions differentiable on a domain and continuous on its closure +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Many theorems in complex analysis assume that a function is complex differentiable on a domain and is continuous on its closure. In this file we define a predicate `diff_cont_on_cl` that expresses this property and prove basic facts about this predicate. diff --git a/src/analysis/calculus/dslope.lean b/src/analysis/calculus/dslope.lean index 810aa968ddc1d..594cc93865f6d 100644 --- a/src/analysis/calculus/dslope.lean +++ b/src/analysis/calculus/dslope.lean @@ -9,6 +9,9 @@ import analysis.calculus.deriv.inv /-! # Slope of a differentiable function +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given a function `f : 𝕜 → E` from a nontrivially normed field to a normed space over this field, `dslope f a b` is defined as `slope f a b = (b - a)⁻¹ • (f b - f a)` for `a ≠ b` and as `deriv f a` for `a = b`. diff --git a/src/analysis/inner_product_space/gram_schmidt_ortho.lean b/src/analysis/inner_product_space/gram_schmidt_ortho.lean index a46b0f5b0a442..ad0ae5dd303c5 100644 --- a/src/analysis/inner_product_space/gram_schmidt_ortho.lean +++ b/src/analysis/inner_product_space/gram_schmidt_ortho.lean @@ -10,6 +10,9 @@ import linear_algebra.matrix.block /-! # Gram-Schmidt Orthogonalization and Orthonormalization +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we introduce Gram-Schmidt Orthogonalization and Orthonormalization. The Gram-Schmidt process takes a set of vectors as input diff --git a/src/category_theory/abelian/ext.lean b/src/category_theory/abelian/ext.lean index 0ee23aa02ea93..c2c6197e87e83 100644 --- a/src/category_theory/abelian/ext.lean +++ b/src/category_theory/abelian/ext.lean @@ -12,6 +12,9 @@ import category_theory.abelian.projective /-! # Ext +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define `Ext R C n : Cᵒᵖ ⥤ C ⥤ Module R` for any `R`-linear abelian category `C` by (left) deriving in the first argument of the bifunctor `(X, Y) ↦ Module.of R (unop X ⟶ Y)`. diff --git a/src/category_theory/abelian/left_derived.lean b/src/category_theory/abelian/left_derived.lean index 8600b4b02535a..1b4a39141973b 100644 --- a/src/category_theory/abelian/left_derived.lean +++ b/src/category_theory/abelian/left_derived.lean @@ -12,6 +12,9 @@ import category_theory.limits.constructions.epi_mono /-! # Zeroth left derived functors +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + If `F : C ⥤ D` is an additive right exact functor between abelian categories, where `C` has enough projectives, we provide the natural isomorphism `F.left_derived 0 ≅ F`. diff --git a/src/category_theory/bicategory/single_obj.lean b/src/category_theory/bicategory/single_obj.lean index 22cb25bbfbe73..8612984699bfb 100644 --- a/src/category_theory/bicategory/single_obj.lean +++ b/src/category_theory/bicategory/single_obj.lean @@ -9,6 +9,9 @@ import category_theory.monoidal.functor /-! # Promoting a monoidal category to a single object bicategory. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A monoidal category can be thought of as a bicategory with a single object. The objects of the monoidal category become the 1-morphisms, diff --git a/src/data/real/golden_ratio.lean b/src/data/real/golden_ratio.lean index 4c0ef8fec16ca..c5963b5fe0114 100644 --- a/src/data/real/golden_ratio.lean +++ b/src/data/real/golden_ratio.lean @@ -13,6 +13,9 @@ import algebra.linear_recurrence /-! # The golden ratio and its conjugate +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the golden ratio `φ := (1 + √5)/2` and its conjugate `ψ := (1 - √5)/2`, which are the two real roots of `X² - X - 1`. diff --git a/src/measure_theory/category/Meas.lean b/src/measure_theory/category/Meas.lean index 07d607a524b27..bfbecf24eff65 100644 --- a/src/measure_theory/category/Meas.lean +++ b/src/measure_theory/category/Meas.lean @@ -11,6 +11,9 @@ import topology.category.Top.basic /-! # The category of measurable spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Measurable spaces and measurable functions form a (concrete) category `Meas`. ## Main definitions diff --git a/src/measure_theory/decomposition/jordan.lean b/src/measure_theory/decomposition/jordan.lean index de697ffb6d442..85d3be72c1027 100644 --- a/src/measure_theory/decomposition/jordan.lean +++ b/src/measure_theory/decomposition/jordan.lean @@ -9,6 +9,9 @@ import measure_theory.measure.mutually_singular /-! # Jordan decomposition +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves the existence and uniqueness of the Jordan decomposition for signed measures. The Jordan decomposition theorem states that, given a signed measure `s`, there exists a unique pair of mutually singular measures `μ` and `ν`, such that `s = μ - ν`. diff --git a/src/measure_theory/decomposition/signed_hahn.lean b/src/measure_theory/decomposition/signed_hahn.lean index 25a9760d1e745..bda6013a8533d 100644 --- a/src/measure_theory/decomposition/signed_hahn.lean +++ b/src/measure_theory/decomposition/signed_hahn.lean @@ -9,6 +9,9 @@ import order.symm_diff /-! # Hahn decomposition +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves the Hahn decomposition theorem (signed version). The Hahn decomposition theorem states that, given a signed measure `s`, there exist complementary, measurable sets `i` and `j`, such that `i` is positive and `j` is negative with respect to `s`; that is, `s` restricted on `i` diff --git a/src/measure_theory/group/action.lean b/src/measure_theory/group/action.lean index 3fa13676beccd..82da7ac4e5014 100644 --- a/src/measure_theory/group/action.lean +++ b/src/measure_theory/group/action.lean @@ -11,6 +11,9 @@ import dynamics.minimal /-! # Measures invariant under group actions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A measure `μ : measure α` is said to be *invariant* under an action of a group `G` if scalar multiplication by `c : G` is a measure preserving map for all `c`. In this file we define a typeclass for measures invariant under action of an (additive or multiplicative) group and prove diff --git a/src/measure_theory/group/measure.lean b/src/measure_theory/group/measure.lean index 8ed1592e22f69..5207c9ebdc00f 100644 --- a/src/measure_theory/group/measure.lean +++ b/src/measure_theory/group/measure.lean @@ -14,6 +14,9 @@ import topology.continuous_function.cocompact_map /-! # Measures on Groups +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We develop some properties of measures on (topological) groups * We define properties on measures: measures that are left or right invariant w.r.t. multiplication. diff --git a/src/number_theory/arithmetic_function.lean b/src/number_theory/arithmetic_function.lean index 6c7ead1903114..066d7936becc0 100644 --- a/src/number_theory/arithmetic_function.lean +++ b/src/number_theory/arithmetic_function.lean @@ -14,6 +14,9 @@ import data.nat.factorization.basic /-! # Arithmetic Functions and Dirichlet Convolution +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines arithmetic functions, which are functions from `ℕ` to a specified type that map 0 to 0. In the literature, they are often instead defined as functions from `ℕ+`. These arithmetic functions are endowed with a multiplication, given by Dirichlet convolution, and pointwise addition, diff --git a/src/number_theory/l_series.lean b/src/number_theory/l_series.lean index 0e6db23043981..13575962b2973 100644 --- a/src/number_theory/l_series.lean +++ b/src/number_theory/l_series.lean @@ -11,6 +11,9 @@ import topology.algebra.infinite_sum.basic /-! # L-series +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given an arithmetic function, we define the corresponding L-series. ## Main Definitions diff --git a/src/number_theory/von_mangoldt.lean b/src/number_theory/von_mangoldt.lean index cbff35d3fb5e2..0d062b3984f10 100644 --- a/src/number_theory/von_mangoldt.lean +++ b/src/number_theory/von_mangoldt.lean @@ -11,6 +11,9 @@ import analysis.special_functions.log.basic /-! # The von Mangoldt Function +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define the von Mangoldt function: the function on natural numbers that returns `log p` if the input can be expressed as `p^k` for a prime `p`. diff --git a/src/topology/algebra/continuous_monoid_hom.lean b/src/topology/algebra/continuous_monoid_hom.lean index f1616da7d2c25..75e8053ff2cac 100644 --- a/src/topology/algebra/continuous_monoid_hom.lean +++ b/src/topology/algebra/continuous_monoid_hom.lean @@ -10,6 +10,9 @@ import topology.continuous_function.algebra # Continuous Monoid Homs +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the space of continuous homomorphisms between two topological groups. ## Main definitions diff --git a/src/topology/instances/discrete.lean b/src/topology/instances/discrete.lean index 3bb76457acdbc..fef680bfeb5f6 100644 --- a/src/topology/instances/discrete.lean +++ b/src/topology/instances/discrete.lean @@ -11,6 +11,9 @@ import topology.metric_space.metrizable_uniformity /-! # Instances related to the discrete topology +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We prove that the discrete topology is * first-countable, * second-countable for an encodable type, diff --git a/src/topology/sheaves/stalks.lean b/src/topology/sheaves/stalks.lean index f6b7277e45807..ea3a7426b187a 100644 --- a/src/topology/sheaves/stalks.lean +++ b/src/topology/sheaves/stalks.lean @@ -17,6 +17,9 @@ import category_theory.sites.pushforward /-! # Stalks +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + For a presheaf `F` on a topological space `X`, valued in some category `C`, the *stalk* of `F` at the point `x : X` is defined as the colimit of the composition of the inclusion of categories `(nhds x)ᵒᵖ ⥤ (opens X)ᵒᵖ` and the functor `F : (opens X)ᵒᵖ ⥤ C`. diff --git a/src/topology/tietze_extension.lean b/src/topology/tietze_extension.lean index 2e43880370bce..5364e367cb487 100644 --- a/src/topology/tietze_extension.lean +++ b/src/topology/tietze_extension.lean @@ -11,6 +11,9 @@ import topology.urysohns_bounded /-! # Tietze extension theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove a few version of the Tietze extension theorem. The theorem says that a continuous function `s → ℝ` defined on a closed set in a normal topological space `Y` can be extended to a continuous function on the whole space. Moreover, if all values of the original From a99f85220eaf38f14f94e04699943e185a5e1d1a Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 31 May 2023 09:30:33 +0000 Subject: [PATCH 1120/1291] feat(linear_algebra/matrix/adjugate): add `det_eq_sum_mul_adjugate_row` (#19117) From lean-matrix-cookbook --- src/data/matrix/notation.lean | 12 +++++++++++ src/linear_algebra/matrix/adjugate.lean | 28 ++++++++++++++++++++++++- 2 files changed, 39 insertions(+), 1 deletion(-) diff --git a/src/data/matrix/notation.lean b/src/data/matrix/notation.lean index 8ff6eefcde35a..f82cd6f7c1009 100644 --- a/src/data/matrix/notation.lean +++ b/src/data/matrix/notation.lean @@ -314,6 +314,18 @@ empty_eq _ submatrix A (vec_cons i row) col = vec_cons (λ j, A i (col j)) (submatrix A row col) := by { ext i j, refine fin.cases _ _ i; simp [submatrix] } +/-- Updating a row then removing it is the same as removing it. -/ +@[simp] lemma submatrix_update_row_succ_above (A : matrix (fin m.succ) n' α) + (v : n' → α) (f : o' → n') (i : fin m.succ) : + (A.update_row i v).submatrix i.succ_above f = A.submatrix i.succ_above f := +ext $ λ r s, (congr_fun (update_row_ne (fin.succ_above_ne i r) : _ = A _) (f s) : _) + +/-- Updating a column then removing it is the same as removing it. -/ +@[simp] lemma submatrix_update_column_succ_above (A : matrix m' (fin n.succ) α) + (v : m' → α) (f : o' → m') (i : fin n.succ) : + (A.update_column i v).submatrix f i.succ_above = A.submatrix f i.succ_above := +ext $ λ r s, update_column_ne (fin.succ_above_ne i s) + end submatrix section vec2_and_vec3 diff --git a/src/linear_algebra/matrix/adjugate.lean b/src/linear_algebra/matrix/adjugate.lean index eb87b9f7f57c4..d2877ef3ef6fa 100644 --- a/src/linear_algebra/matrix/adjugate.lean +++ b/src/linear_algebra/matrix/adjugate.lean @@ -340,7 +340,6 @@ lemma _root_.alg_hom.map_adjugate {R A B : Type*} [comm_semiring R] [comm_ring A (M : matrix n n A) : f.map_matrix M.adjugate = matrix.adjugate (f.map_matrix M) := f.to_ring_hom.map_adjugate _ - lemma det_adjugate (A : matrix n n α) : (adjugate A).det = A.det ^ (fintype.card n - 1) := begin -- get rid of the `- 1` @@ -387,6 +386,33 @@ end adjugate !![a, b; c, d] = !![d, -b; -c, a] := adjugate_fin_two _ +lemma adjugate_fin_succ_eq_det_submatrix {n : ℕ} (A : matrix (fin n.succ) (fin n.succ) α) (i j) : + adjugate A i j = (-1)^(j + i : ℕ) * det (A.submatrix j.succ_above i.succ_above) := +begin + simp_rw [adjugate_apply, det_succ_row _ j, update_row_self, submatrix_update_row_succ_above], + rw [fintype.sum_eq_single i (λ h hjk, _), pi.single_eq_same, mul_one], + rw [pi.single_eq_of_ne hjk, mul_zero, zero_mul], +end + +lemma det_eq_sum_mul_adjugate_row (A : matrix n n α) (i : n) : + det A = ∑ j : n, A i j * adjugate A j i := +begin + haveI : nonempty n := ⟨i⟩, + obtain ⟨n', hn'⟩ := nat.exists_eq_succ_of_ne_zero (fintype.card_ne_zero : fintype.card n ≠ 0), + obtain ⟨e⟩ := fintype.trunc_equiv_fin_of_card_eq hn', + let A' := reindex e e A, + suffices : det A' = ∑ j : fin n'.succ, A' (e i) j * adjugate A' j (e i), + { simp_rw [A', det_reindex_self, adjugate_reindex, reindex_apply, submatrix_apply, ←e.sum_comp, + equiv.symm_apply_apply] at this, + exact this }, + rw det_succ_row A' (e i), + simp_rw [mul_assoc, mul_left_comm _ (A' _ _), ←adjugate_fin_succ_eq_det_submatrix], +end + +lemma det_eq_sum_mul_adjugate_col (A : matrix n n α) (j : n) : + det A = ∑ i : n, A i j * adjugate A j i := +by simpa only [det_transpose, ←adjugate_transpose] using det_eq_sum_mul_adjugate_row Aᵀ j + lemma adjugate_conj_transpose [star_ring α] (A : matrix n n α) : A.adjugateᴴ = adjugate (Aᴴ) := begin dsimp only [conj_transpose], From 00abe0695d8767201e6d008afa22393978bb324d Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Wed, 31 May 2023 12:41:58 +0000 Subject: [PATCH 1121/1291] feat(probability/kernel/cond_distrib): regular conditional probability distributions (#19090) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We define the regular conditional probability distribution `cond_distrib Y X μ` of `Y : α → Ω` given `X : α → β`, where `Ω` is a standard Borel space. This is a `kernel β Ω` such that for almost all `a`, for all measurable set `s`, `cond_distrib Y X μ (X a) s` is equal to the conditional expectation `μ⟦Y ⁻¹' s | mβ.comap X⟧` evaluated at `a`. Also define the above notation for the conditional expectation of the indicator of a set. Co-authored-by: Rémy Degenne --- .../constructions/prod/basic.lean | 30 ++ .../conditional_expectation/basic.lean | 8 + src/probability/kernel/cond_distrib.lean | 315 ++++++++++++++++++ src/probability/kernel/condexp.lean | 175 ++++++++++ src/probability/notation.lean | 6 + 5 files changed, 534 insertions(+) create mode 100644 src/probability/kernel/cond_distrib.lean create mode 100644 src/probability/kernel/condexp.lean diff --git a/src/measure_theory/constructions/prod/basic.lean b/src/measure_theory/constructions/prod/basic.lean index 23adab12a0893..7aa1cd9ca24b2 100644 --- a/src/measure_theory/constructions/prod/basic.lean +++ b/src/measure_theory/constructions/prod/basic.lean @@ -749,6 +749,21 @@ instance [is_finite_measure ρ] : is_finite_measure ρ.fst := by { rw fst, apply instance [is_probability_measure ρ] : is_probability_measure ρ.fst := { measure_univ := by { rw fst_univ, exact measure_univ, } } +lemma fst_map_prod_mk₀ {X : α → β} {Y : α → γ} {μ : measure α} + (hX : ae_measurable X μ) (hY : ae_measurable Y μ) : + (μ.map (λ a, (X a, Y a))).fst = μ.map X := +begin + ext1 s hs, + rw [measure.fst_apply hs, measure.map_apply_of_ae_measurable (hX.prod_mk hY) (measurable_fst hs), + measure.map_apply_of_ae_measurable hX hs, ← prod_univ, mk_preimage_prod, preimage_univ, + inter_univ], +end + +lemma fst_map_prod_mk {X : α → β} {Y : α → γ} {μ : measure α} + (hX : measurable X) (hY : measurable Y) : + (μ.map (λ a, (X a, Y a))).fst = μ.map X := +fst_map_prod_mk₀ hX.ae_measurable hY.ae_measurable + /-- Marginal measure on `β` obtained from a measure on `ρ` `α × β`, defined by `ρ.map prod.snd`. -/ noncomputable def snd (ρ : measure (α × β)) : measure β := ρ.map prod.snd @@ -764,6 +779,21 @@ instance [is_finite_measure ρ] : is_finite_measure ρ.snd := by { rw snd, apply instance [is_probability_measure ρ] : is_probability_measure ρ.snd := { measure_univ := by { rw snd_univ, exact measure_univ, } } +lemma snd_map_prod_mk₀ {X : α → β} {Y : α → γ} {μ : measure α} + (hX : ae_measurable X μ) (hY : ae_measurable Y μ) : + (μ.map (λ a, (X a, Y a))).snd = μ.map Y := +begin + ext1 s hs, + rw [measure.snd_apply hs, measure.map_apply_of_ae_measurable (hX.prod_mk hY) (measurable_snd hs), + measure.map_apply_of_ae_measurable hY hs, ← univ_prod, mk_preimage_prod, preimage_univ, + univ_inter], +end + +lemma snd_map_prod_mk {X : α → β} {Y : α → γ} {μ : measure α} + (hX : measurable X) (hY : measurable Y) : + (μ.map (λ a, (X a, Y a))).snd = μ.map Y := +snd_map_prod_mk₀ hX.ae_measurable hY.ae_measurable + end measure diff --git a/src/measure_theory/function/conditional_expectation/basic.lean b/src/measure_theory/function/conditional_expectation/basic.lean index 6921a058ecaf7..1bd8bbd85c74e 100644 --- a/src/measure_theory/function/conditional_expectation/basic.lean +++ b/src/measure_theory/function/conditional_expectation/basic.lean @@ -187,6 +187,14 @@ lemma ae_eq_trim_iff_of_ae_strongly_measurable' {α β} [topological_space β] [ ⟨λ h, hfm.ae_eq_mk.trans (h.trans hgm.ae_eq_mk.symm), λ h, hfm.ae_eq_mk.symm.trans (h.trans hgm.ae_eq_mk)⟩ +lemma ae_strongly_measurable.comp_ae_measurable' + {α β γ : Type*} [topological_space β] {mα : measurable_space α} {mγ : measurable_space γ} + {f : α → β} {μ : measure γ} {g : γ → α} + (hf : ae_strongly_measurable f (μ.map g)) (hg : ae_measurable g μ) : + ae_strongly_measurable' (mα.comap g) (f ∘ g) μ := +⟨(hf.mk f) ∘ g, hf.strongly_measurable_mk.comp_measurable (measurable_iff_comap_le.mpr le_rfl), + ae_eq_comp hg hf.ae_eq_mk⟩ + /-- If the restriction to a set `s` of a σ-algebra `m` is included in the restriction to `s` of another σ-algebra `m₂` (hypothesis `hs`), the set `s` is `m` measurable and a function `f` almost everywhere supported on `s` is `m`-ae-strongly-measurable, then `f` is also diff --git a/src/probability/kernel/cond_distrib.lean b/src/probability/kernel/cond_distrib.lean new file mode 100644 index 0000000000000..6da8554ccaf4b --- /dev/null +++ b/src/probability/kernel/cond_distrib.lean @@ -0,0 +1,315 @@ +/- +Copyright (c) 2023 Rémy Degenne. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Rémy Degenne +-/ +import probability.kernel.disintegration +import probability.notation + +/-! +# Regular conditional probability distribution + +We define the regular conditional probability distribution of `Y : α → Ω` given `X : α → β`, where +`Ω` is a standard Borel space. This is a `kernel β Ω` such that for almost all `a`, `cond_distrib` +evaluated at `X a` and a measurable set `s` is equal to the conditional expectation +`μ⟦Y ⁻¹' s | mβ.comap X⟧` evaluated at `a`. + +`μ⟦Y ⁻¹' s | mβ.comap X⟧` maps a measurable set `s` to a function `α → ℝ≥0∞`, and for all `s` that +map is unique up to a `μ`-null set. For all `a`, the map from sets to `ℝ≥0∞` that we obtain that way +verifies some of the properties of a measure, but in general the fact that the `μ`-null set depends +on `s` can prevent us from finding versions of the conditional expectation that combine into a true +measure. The standard Borel space assumption on `Ω` allows us to do so. + +The case `Y = X = id` is developed in more detail in `probability/kernel/condexp.lean`: here `X` is +understood as a map from `Ω` with a sub-σ-algebra to `Ω` with its default σ-algebra and the +conditional distribution defines a kernel associated with the conditional expectation with respect +to `m`. + +## Main definitions + +* `cond_distrib Y X μ`: regular conditional probability distribution of `Y : α → Ω` given + `X : α → β`, where `Ω` is a standard Borel space. + +## Main statements + +* `cond_distrib_ae_eq_condexp`: for almost all `a`, `cond_distrib` evaluated at `X a` and a + measurable set `s` is equal to the conditional expectation `μ⟦Y ⁻¹' s | mβ.comap X⟧ a`. +* `condexp_prod_ae_eq_integral_cond_distrib`: the conditional expectation + `μ[(λ a, f (X a, Y a)) | X ; mβ]` is almost everywhere equal to the integral + `∫ y, f (X a, y) ∂(cond_distrib Y X μ (X a))`. + +-/ + +open measure_theory set filter topological_space + +open_locale ennreal measure_theory probability_theory + +namespace probability_theory + +variables {α β Ω F : Type*} + [topological_space Ω] [measurable_space Ω] [polish_space Ω] [borel_space Ω] [nonempty Ω] + [normed_add_comm_group F] + {mα : measurable_space α} {μ : measure α} [is_finite_measure μ] {X : α → β} {Y : α → Ω} + +/-- **Regular conditional probability distribution**: kernel associated with the conditional +expectation of `Y` given `X`. +For almost all `a`, `cond_distrib Y X μ` evaluated at `X a` and a measurable set `s` is equal to +the conditional expectation `μ⟦Y ⁻¹' s | mβ.comap X⟧ a`. It also satisfies the equality +`μ[(λ a, f (X a, Y a)) | mβ.comap X] =ᵐ[μ] λ a, ∫ y, f (X a, y) ∂(cond_distrib Y X μ (X a))` for +all integrable functions `f`. -/ +@[irreducible] noncomputable +def cond_distrib {mα : measurable_space α} [measurable_space β] + (Y : α → Ω) (X : α → β) (μ : measure α) [is_finite_measure μ] : + kernel β Ω := +(μ.map (λ a, (X a, Y a))).cond_kernel + +instance [measurable_space β] : is_markov_kernel (cond_distrib Y X μ) := +by { rw cond_distrib, apply_instance, } + +variables {mβ : measurable_space β} {s : set Ω} {t : set β} {f : β × Ω → F} +include mβ + +section measurability + +lemma measurable_cond_distrib (hs : measurable_set s) : + measurable[mβ.comap X] (λ a, cond_distrib Y X μ (X a) s) := +(kernel.measurable_coe _ hs).comp (measurable.of_comap_le le_rfl) + +lemma _root_.measure_theory.ae_strongly_measurable.ae_integrable_cond_distrib_map_iff + (hX : ae_measurable X μ) (hY : ae_measurable Y μ) + (hf : ae_strongly_measurable f (μ.map (λ a, (X a, Y a)))) : + (∀ᵐ a ∂(μ.map X), integrable (λ ω, f (a, ω)) (cond_distrib Y X μ a)) + ∧ integrable (λ a, ∫ ω, ‖f (a, ω)‖ ∂(cond_distrib Y X μ a)) (μ.map X) + ↔ integrable f (μ.map (λ a, (X a, Y a))) := +by rw [cond_distrib, ← hf.ae_integrable_cond_kernel_iff, measure.fst_map_prod_mk₀ hX hY] + +variables [normed_space ℝ F] [complete_space F] + +lemma _root_.measure_theory.ae_strongly_measurable.integral_cond_distrib_map + (hX : ae_measurable X μ) (hY : ae_measurable Y μ) + (hf : ae_strongly_measurable f (μ.map (λ a, (X a, Y a)))) : + ae_strongly_measurable (λ x, ∫ y, f (x, y) ∂(cond_distrib Y X μ x)) (μ.map X) := +by { rw [← measure.fst_map_prod_mk₀ hX hY, cond_distrib], exact hf.integral_cond_kernel, } + +lemma _root_.measure_theory.ae_strongly_measurable.integral_cond_distrib + (hX : ae_measurable X μ) (hY : ae_measurable Y μ) + (hf : ae_strongly_measurable f (μ.map (λ a, (X a, Y a)))) : + ae_strongly_measurable (λ a, ∫ y, f (X a, y) ∂(cond_distrib Y X μ (X a))) μ := +(hf.integral_cond_distrib_map hX hY).comp_ae_measurable hX + +lemma ae_strongly_measurable'_integral_cond_distrib + (hX : ae_measurable X μ) (hY : ae_measurable Y μ) + (hf : ae_strongly_measurable f (μ.map (λ a, (X a, Y a)))) : + ae_strongly_measurable' (mβ.comap X) (λ a, ∫ y, f (X a, y) ∂(cond_distrib Y X μ (X a))) μ := +(hf.integral_cond_distrib_map hX hY).comp_ae_measurable' hX + +end measurability + +section integrability + +lemma integrable_to_real_cond_distrib (hX : ae_measurable X μ) (hs : measurable_set s) : + integrable (λ a, (cond_distrib Y X μ (X a) s).to_real) μ := +begin + refine integrable_to_real_of_lintegral_ne_top _ _, + { exact measurable.comp_ae_measurable (kernel.measurable_coe _ hs) hX, }, + { refine ne_of_lt _, + calc ∫⁻ a, cond_distrib Y X μ (X a) s ∂μ + ≤ ∫⁻ a, 1 ∂μ : lintegral_mono (λ a, prob_le_one) + ... = μ univ : lintegral_one + ... < ∞ : measure_lt_top _ _, }, +end + +lemma _root_.measure_theory.integrable.cond_distrib_ae_map + (hX : ae_measurable X μ) (hY : ae_measurable Y μ) + (hf_int : integrable f (μ.map (λ a, (X a, Y a)))) : + ∀ᵐ b ∂(μ.map X), integrable (λ ω, f (b, ω)) (cond_distrib Y X μ b) := +by { rw [cond_distrib, ← measure.fst_map_prod_mk₀ hX hY], exact hf_int.cond_kernel_ae, } + +lemma _root_.measure_theory.integrable.cond_distrib_ae + (hX : ae_measurable X μ) (hY : ae_measurable Y μ) + (hf_int : integrable f (μ.map (λ a, (X a, Y a)))) : + ∀ᵐ a ∂μ, integrable (λ ω, f (X a, ω)) (cond_distrib Y X μ (X a)) := +ae_of_ae_map hX (hf_int.cond_distrib_ae_map hX hY) + +lemma _root_.measure_theory.integrable.integral_norm_cond_distrib_map + (hX : ae_measurable X μ) (hY : ae_measurable Y μ) + (hf_int : integrable f (μ.map (λ a, (X a, Y a)))) : + integrable (λ x, ∫ y, ‖f (x, y)‖ ∂(cond_distrib Y X μ x)) (μ.map X) := +by { rw [cond_distrib, ← measure.fst_map_prod_mk₀ hX hY], exact hf_int.integral_norm_cond_kernel, } + +lemma _root_.measure_theory.integrable.integral_norm_cond_distrib + (hX : ae_measurable X μ) (hY : ae_measurable Y μ) + (hf_int : integrable f (μ.map (λ a, (X a, Y a)))) : + integrable (λ a, ∫ y, ‖f (X a, y)‖ ∂(cond_distrib Y X μ (X a))) μ := +(hf_int.integral_norm_cond_distrib_map hX hY).comp_ae_measurable hX + +variables [normed_space ℝ F] [complete_space F] + +lemma _root_.measure_theory.integrable.norm_integral_cond_distrib_map + (hX : ae_measurable X μ) (hY : ae_measurable Y μ) + (hf_int : integrable f (μ.map (λ a, (X a, Y a)))) : + integrable (λ x, ‖∫ y, f (x, y) ∂(cond_distrib Y X μ x)‖) (μ.map X) := +by { rw [cond_distrib, ← measure.fst_map_prod_mk₀ hX hY], exact hf_int.norm_integral_cond_kernel, } + +lemma _root_.measure_theory.integrable.norm_integral_cond_distrib + (hX : ae_measurable X μ) (hY : ae_measurable Y μ) + (hf_int : integrable f (μ.map (λ a, (X a, Y a)))) : + integrable (λ a, ‖∫ y, f (X a, y) ∂(cond_distrib Y X μ (X a))‖) μ := +(hf_int.norm_integral_cond_distrib_map hX hY).comp_ae_measurable hX + +lemma _root_.measure_theory.integrable.integral_cond_distrib_map + (hX : ae_measurable X μ) (hY : ae_measurable Y μ) + (hf_int : integrable f (μ.map (λ a, (X a, Y a)))) : + integrable (λ x, ∫ y, f (x, y) ∂(cond_distrib Y X μ x)) (μ.map X) := +(integrable_norm_iff (hf_int.1.integral_cond_distrib_map hX hY)).mp + (hf_int.norm_integral_cond_distrib_map hX hY) + +lemma _root_.measure_theory.integrable.integral_cond_distrib + (hX : ae_measurable X μ) (hY : ae_measurable Y μ) + (hf_int : integrable f (μ.map (λ a, (X a, Y a)))) : + integrable (λ a, ∫ y, f (X a, y) ∂(cond_distrib Y X μ (X a))) μ := +(hf_int.integral_cond_distrib_map hX hY).comp_ae_measurable hX + +end integrability + +lemma set_lintegral_preimage_cond_distrib (hX : measurable X) (hY : ae_measurable Y μ) + (hs : measurable_set s) (ht : measurable_set t) : + ∫⁻ a in X ⁻¹' t, cond_distrib Y X μ (X a) s ∂μ = μ (X ⁻¹' t ∩ Y ⁻¹' s) := +by rw [lintegral_comp (kernel.measurable_coe _ hs) hX, cond_distrib, + ← measure.restrict_map hX ht, ← measure.fst_map_prod_mk₀ hX.ae_measurable hY, + set_lintegral_cond_kernel_eq_measure_prod _ ht hs, + measure.map_apply_of_ae_measurable (hX.ae_measurable.prod_mk hY) (ht.prod hs), + mk_preimage_prod] + +lemma set_lintegral_cond_distrib_of_measurable_set (hX : measurable X) (hY : ae_measurable Y μ) + (hs : measurable_set s) {t : set α} (ht : measurable_set[mβ.comap X] t) : + ∫⁻ a in t, cond_distrib Y X μ (X a) s ∂μ = μ (t ∩ Y ⁻¹' s) := +by { obtain ⟨t', ht', rfl⟩ := ht, rw set_lintegral_preimage_cond_distrib hX hY hs ht', } + +/-- For almost every `a : α`, the `cond_distrib Y X μ` kernel applied to `X a` and a measurable set +`s` is equal to the conditional expectation of the indicator of `Y ⁻¹' s`. -/ +lemma cond_distrib_ae_eq_condexp (hX : measurable X) (hY : measurable Y) (hs : measurable_set s) : + (λ a, (cond_distrib Y X μ (X a) s).to_real) =ᵐ[μ] μ⟦Y ⁻¹' s | mβ.comap X⟧ := +begin + refine ae_eq_condexp_of_forall_set_integral_eq hX.comap_le _ _ _ _, + { exact (integrable_const _).indicator (hY hs), }, + { exact λ t ht _, (integrable_to_real_cond_distrib hX.ae_measurable hs).integrable_on, }, + { intros t ht _, + rw [integral_to_real ((measurable_cond_distrib hs).mono hX.comap_le le_rfl).ae_measurable + (eventually_of_forall (λ ω, measure_lt_top (cond_distrib Y X μ (X ω)) _)), + integral_indicator_const _ (hY hs), measure.restrict_apply (hY hs), smul_eq_mul, mul_one, + inter_comm, set_lintegral_cond_distrib_of_measurable_set hX hY.ae_measurable hs ht], }, + { refine (measurable.strongly_measurable _).ae_strongly_measurable', + exact @measurable.ennreal_to_real _ (mβ.comap X) _ (measurable_cond_distrib hs), }, +end + +/-- The conditional expectation of a function `f` of the product `(X, Y)` is almost everywhere equal +to the integral of `y ↦ f(X, y)` against the `cond_distrib` kernel. -/ +lemma condexp_prod_ae_eq_integral_cond_distrib' [normed_space ℝ F] [complete_space F] + (hX : measurable X) (hY : ae_measurable Y μ) + (hf_int : integrable f (μ.map (λ a, (X a, Y a)))) : + μ[(λ a, f (X a, Y a)) | mβ.comap X] =ᵐ[μ] λ a, ∫ y, f (X a, y) ∂(cond_distrib Y X μ (X a)) := +begin + have hf_int' : integrable (λ a, f (X a, Y a)) μ, + { exact (integrable_map_measure hf_int.1 (hX.ae_measurable.prod_mk hY)).mp hf_int, }, + refine (ae_eq_condexp_of_forall_set_integral_eq hX.comap_le hf_int' (λ s hs hμs, _) _ _).symm, + { exact (hf_int.integral_cond_distrib hX.ae_measurable hY).integrable_on, }, + { rintros s ⟨t, ht, rfl⟩ _, + change ∫ a in X ⁻¹' t, ((λ x', ∫ y, f (x', y) ∂(cond_distrib Y X μ) x') ∘ X) a ∂μ + = ∫ a in X ⁻¹' t, f (X a, Y a) ∂μ, + rw ← integral_map hX.ae_measurable, + swap, + { rw ← measure.restrict_map hX ht, + exact (hf_int.1.integral_cond_distrib_map hX.ae_measurable hY).restrict, }, + rw [← measure.restrict_map hX ht, ← measure.fst_map_prod_mk₀ hX.ae_measurable hY, cond_distrib, + set_integral_cond_kernel_univ_right ht hf_int.integrable_on, + set_integral_map (ht.prod measurable_set.univ) hf_int.1 (hX.ae_measurable.prod_mk hY), + mk_preimage_prod, preimage_univ, inter_univ], }, + { exact ae_strongly_measurable'_integral_cond_distrib hX.ae_measurable hY hf_int.1, }, +end + +/-- The conditional expectation of a function `f` of the product `(X, Y)` is almost everywhere equal +to the integral of `y ↦ f(X, y)` against the `cond_distrib` kernel. -/ +lemma condexp_prod_ae_eq_integral_cond_distrib₀ [normed_space ℝ F] [complete_space F] + (hX : measurable X) (hY : ae_measurable Y μ) + (hf : ae_strongly_measurable f (μ.map (λ a, (X a, Y a)))) + (hf_int : integrable (λ a, f (X a, Y a)) μ) : + μ[(λ a, f (X a, Y a)) | mβ.comap X] =ᵐ[μ] λ a, ∫ y, f (X a, y) ∂(cond_distrib Y X μ (X a)) := +begin + have hf_int' : integrable f (μ.map (λ a, (X a, Y a))), + { rwa integrable_map_measure hf (hX.ae_measurable.prod_mk hY), }, + exact condexp_prod_ae_eq_integral_cond_distrib' hX hY hf_int', +end + +/-- The conditional expectation of a function `f` of the product `(X, Y)` is almost everywhere equal +to the integral of `y ↦ f(X, y)` against the `cond_distrib` kernel. -/ +lemma condexp_prod_ae_eq_integral_cond_distrib [normed_space ℝ F] [complete_space F] + (hX : measurable X) (hY : ae_measurable Y μ) + (hf : strongly_measurable f) (hf_int : integrable (λ a, f (X a, Y a)) μ) : + μ[(λ a, f (X a, Y a)) | mβ.comap X] =ᵐ[μ] λ a, ∫ y, f (X a, y) ∂(cond_distrib Y X μ (X a)) := +begin + have hf_int' : integrable f (μ.map (λ a, (X a, Y a))), + { rwa integrable_map_measure hf.ae_strongly_measurable (hX.ae_measurable.prod_mk hY), }, + exact condexp_prod_ae_eq_integral_cond_distrib' hX hY hf_int', +end + +lemma condexp_ae_eq_integral_cond_distrib [normed_space ℝ F] [complete_space F] + (hX : measurable X) (hY : ae_measurable Y μ) + {f : Ω → F} (hf : strongly_measurable f) (hf_int : integrable (λ a, f (Y a)) μ) : + μ[(λ a, f (Y a)) | mβ.comap X] =ᵐ[μ] λ a, ∫ y, f y ∂(cond_distrib Y X μ (X a)) := +condexp_prod_ae_eq_integral_cond_distrib hX hY (hf.comp_measurable measurable_snd) hf_int + +/-- The conditional expectation of `Y` given `X` is almost everywhere equal to the integral +`∫ y, y ∂(cond_distrib Y X μ (X a))`. -/ +lemma condexp_ae_eq_integral_cond_distrib' {Ω} [normed_add_comm_group Ω] [normed_space ℝ Ω] + [complete_space Ω] [measurable_space Ω] [borel_space Ω] [second_countable_topology Ω] {Y : α → Ω} + (hX : measurable X) (hY_int : integrable Y μ) : + μ[Y | mβ.comap X] =ᵐ[μ] λ a, ∫ y, y ∂(cond_distrib Y X μ (X a)) := +condexp_ae_eq_integral_cond_distrib hX hY_int.1.ae_measurable strongly_measurable_id hY_int + +lemma _root_.measure_theory.ae_strongly_measurable.comp_snd_map_prod_mk + {Ω F} {mΩ : measurable_space Ω} {X : Ω → β} {μ : measure Ω} + [topological_space F] (hX : measurable X) {f : Ω → F} (hf : ae_strongly_measurable f μ) : + ae_strongly_measurable (λ x : β × Ω, f x.2) (μ.map (λ ω, (X ω, ω))) := +begin + refine ⟨λ x, hf.mk f x.2, hf.strongly_measurable_mk.comp_measurable measurable_snd, _⟩, + suffices h : measure.quasi_measure_preserving prod.snd (μ.map (λ ω, (X ω, ω))) μ, + { exact measure.quasi_measure_preserving.ae_eq h hf.ae_eq_mk, }, + refine ⟨measurable_snd, measure.absolutely_continuous.mk (λ s hs hμs, _)⟩, + rw measure.map_apply _ hs, + swap, { exact measurable_snd, }, + rw measure.map_apply, + { rw [← univ_prod, mk_preimage_prod, preimage_univ, univ_inter, preimage_id'], + exact hμs, }, + { exact hX.prod_mk measurable_id, }, + { exact measurable_snd hs, }, +end + +lemma _root_.measure_theory.integrable.comp_snd_map_prod_mk {Ω} {mΩ : measurable_space Ω} + {X : Ω → β} {μ : measure Ω} (hX : measurable X) {f : Ω → F} (hf_int : integrable f μ) : + integrable (λ x : β × Ω, f x.2) (μ.map (λ ω, (X ω, ω))) := +begin + have hf := hf_int.1.comp_snd_map_prod_mk hX, + refine ⟨hf, _⟩, + rw [has_finite_integral, lintegral_map' hf.ennnorm (hX.prod_mk measurable_id).ae_measurable], + exact hf_int.2, +end + +lemma ae_strongly_measurable_comp_snd_map_prod_mk_iff {Ω F} {mΩ : measurable_space Ω} + [topological_space F] {X : Ω → β} {μ : measure Ω} (hX : measurable X) {f : Ω → F} : + ae_strongly_measurable (λ x : β × Ω, f x.2) (μ.map (λ ω, (X ω, ω))) + ↔ ae_strongly_measurable f μ := +⟨λ h, h.comp_measurable (hX.prod_mk measurable_id), λ h, h.comp_snd_map_prod_mk hX⟩ + +lemma integrable_comp_snd_map_prod_mk_iff {Ω} {mΩ : measurable_space Ω} {X : Ω → β} {μ : measure Ω} + (hX : measurable X) {f : Ω → F} : + integrable (λ x : β × Ω, f x.2) (μ.map (λ ω, (X ω, ω))) ↔ integrable f μ := +⟨λ h, h.comp_measurable (hX.prod_mk measurable_id), λ h, h.comp_snd_map_prod_mk hX⟩ + +lemma condexp_ae_eq_integral_cond_distrib_id [normed_space ℝ F] [complete_space F] + {X : Ω → β} {μ : measure Ω} [is_finite_measure μ] + (hX : measurable X) {f : Ω → F} (hf_int : integrable f μ) : + μ[f | mβ.comap X] =ᵐ[μ] λ a, ∫ y, f y ∂(cond_distrib id X μ (X a)) := +condexp_prod_ae_eq_integral_cond_distrib' hX ae_measurable_id (hf_int.comp_snd_map_prod_mk hX) + +end probability_theory diff --git a/src/probability/kernel/condexp.lean b/src/probability/kernel/condexp.lean new file mode 100644 index 0000000000000..d860a20e4c3ac --- /dev/null +++ b/src/probability/kernel/condexp.lean @@ -0,0 +1,175 @@ +/- +Copyright (c) 2023 Rémy Degenne. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Rémy Degenne +-/ +import probability.kernel.cond_distrib + +/-! +# Kernel associated with a conditional expectation + +We define `condexp_kernel μ m`, a kernel from `Ω` to `Ω` such that for all integrable functions `f`, +`μ[f | m] =ᵐ[μ] λ ω, ∫ y, f y ∂(condexp_kernel μ m ω)`. + +This kernel is defined if `Ω` is a standard Borel space. In general, `μ⟦s | m⟧` maps a measurable +set `s` to a function `Ω → ℝ≥0∞`, and for all `s` that map is unique up to a `μ`-null set. For all +`a`, the map from sets to `ℝ≥0∞` that we obtain that way verifies some of the properties of a +measure, but the fact that the `μ`-null set depends on `s` can prevent us from finding versions of +the conditional expectation that combine into a true measure. The standard Borel space assumption +on `Ω` allows us to do so. + +## Main definitions + +* `condexp_kernel μ m`: kernel such that `μ[f | m] =ᵐ[μ] λ ω, ∫ y, f y ∂(condexp_kernel μ m ω)`. + +## Main statements + +* `condexp_ae_eq_integral_condexp_kernel`: `μ[f | m] =ᵐ[μ] λ ω, ∫ y, f y ∂(condexp_kernel μ m ω)`. + +-/ + +open measure_theory set filter topological_space + +open_locale ennreal measure_theory probability_theory + +namespace probability_theory + +section aux_lemmas + +variables {Ω F : Type*} {m mΩ : measurable_space Ω} {μ : measure Ω} {f : Ω → F} + +-- todo after the port: move to measure_theory/measurable_space, after measurable.mono +lemma measurable_id'' (hm : m ≤ mΩ) : + @measurable Ω Ω mΩ m id := +measurable_id.mono le_rfl hm + +-- todo after the port: move to measure_theory/measurable_space, after measurable.mono +lemma ae_measurable_id'' (μ : measure Ω) (hm : m ≤ mΩ) : + @ae_measurable Ω Ω m mΩ id μ := +@measurable.ae_measurable Ω Ω mΩ m id μ (measurable_id'' hm) + +lemma _root_.measure_theory.ae_strongly_measurable.comp_snd_map_prod_id [topological_space F] + (hm : m ≤ mΩ) (hf : ae_strongly_measurable f μ) : + ae_strongly_measurable (λ x : Ω × Ω, f x.2) + (@measure.map Ω (Ω × Ω) (m.prod mΩ) mΩ (λ ω, (id ω, id ω)) μ) := +begin + rw ← ae_strongly_measurable_comp_snd_map_prod_mk_iff (measurable_id'' hm) at hf, + simp_rw [id.def] at hf ⊢, + exact hf, +end + +lemma _root_.measure_theory.integrable.comp_snd_map_prod_id [normed_add_comm_group F] + (hm : m ≤ mΩ) (hf : integrable f μ) : + integrable (λ x : Ω × Ω, f x.2) + (@measure.map Ω (Ω × Ω) (m.prod mΩ) mΩ (λ ω, (id ω, id ω)) μ) := +begin + rw ← integrable_comp_snd_map_prod_mk_iff (measurable_id'' hm) at hf, + simp_rw [id.def] at hf ⊢, + exact hf, +end + +end aux_lemmas + +variables {Ω F : Type*} [topological_space Ω] {m : measurable_space Ω} [mΩ : measurable_space Ω] + [polish_space Ω] [borel_space Ω] [nonempty Ω] + {μ : measure Ω} [is_finite_measure μ] + [normed_add_comm_group F] {f : Ω → F} + +/-- Kernel associated with the conditional expectation with respect to a σ-algebra. It satisfies +`μ[f | m] =ᵐ[μ] λ ω, ∫ y, f y ∂(condexp_kernel μ m ω)`. +It is defined as the conditional distribution of the identity given the identity, where the second +identity is understood as a map from `Ω` with the σ-algebra `mΩ` to `Ω` with σ-algebra `m`. -/ +@[irreducible] noncomputable +def condexp_kernel (μ : measure Ω) [is_finite_measure μ] (m : measurable_space Ω) : + @kernel Ω Ω m mΩ := +@cond_distrib Ω Ω Ω _ mΩ _ _ _ mΩ m id id μ _ + +section measurability + +lemma measurable_condexp_kernel {s : set Ω} (hs : measurable_set s) : + measurable[m] (λ ω, condexp_kernel μ m ω s) := +by { rw condexp_kernel, convert measurable_cond_distrib hs, rw measurable_space.comap_id, } + +lemma _root_.measure_theory.ae_strongly_measurable.integral_condexp_kernel + [normed_space ℝ F] [complete_space F] + (hm : m ≤ mΩ) (hf : ae_strongly_measurable f μ) : + ae_strongly_measurable (λ ω, ∫ y, f y ∂(condexp_kernel μ m ω)) μ := +begin + rw condexp_kernel, + exact ae_strongly_measurable.integral_cond_distrib + (ae_measurable_id'' μ hm) ae_measurable_id (hf.comp_snd_map_prod_id hm), +end + +lemma ae_strongly_measurable'_integral_condexp_kernel [normed_space ℝ F] [complete_space F] + (hm : m ≤ mΩ) (hf : ae_strongly_measurable f μ) : + ae_strongly_measurable' m (λ ω, ∫ y, f y ∂(condexp_kernel μ m ω)) μ := +begin + rw condexp_kernel, + have h := ae_strongly_measurable'_integral_cond_distrib + (ae_measurable_id'' μ hm) ae_measurable_id (hf.comp_snd_map_prod_id hm), + rwa measurable_space.comap_id at h, +end + +end measurability + +section integrability + +lemma _root_.measure_theory.integrable.condexp_kernel_ae + (hm : m ≤ mΩ) (hf_int : integrable f μ) : + ∀ᵐ ω ∂μ, integrable f (condexp_kernel μ m ω) := +begin + rw condexp_kernel, + exact integrable.cond_distrib_ae (ae_measurable_id'' μ hm) + ae_measurable_id (hf_int.comp_snd_map_prod_id hm), +end + +lemma _root_.measure_theory.integrable.integral_norm_condexp_kernel + (hm : m ≤ mΩ) (hf_int : integrable f μ) : + integrable (λ ω, ∫ y, ‖f y‖ ∂(condexp_kernel μ m ω)) μ := +begin + rw condexp_kernel, + exact integrable.integral_norm_cond_distrib (ae_measurable_id'' μ hm) + ae_measurable_id (hf_int.comp_snd_map_prod_id hm), +end + +lemma _root_.measure_theory.integrable.norm_integral_condexp_kernel + [normed_space ℝ F] [complete_space F] + (hm : m ≤ mΩ) (hf_int : integrable f μ) : + integrable (λ ω, ‖∫ y, f y ∂(condexp_kernel μ m ω)‖) μ := +begin + rw condexp_kernel, + exact integrable.norm_integral_cond_distrib (ae_measurable_id'' μ hm) + ae_measurable_id (hf_int.comp_snd_map_prod_id hm), +end + +lemma _root_.measure_theory.integrable.integral_condexp_kernel [normed_space ℝ F] [complete_space F] + (hm : m ≤ mΩ) (hf_int : integrable f μ) : + integrable (λ ω, ∫ y, f y ∂(condexp_kernel μ m ω)) μ := +begin + rw condexp_kernel, + exact integrable.integral_cond_distrib (ae_measurable_id'' μ hm) + ae_measurable_id (hf_int.comp_snd_map_prod_id hm), +end + +lemma integrable_to_real_condexp_kernel (hm : m ≤ mΩ) {s : set Ω} (hs : measurable_set s) : + integrable (λ ω, (condexp_kernel μ m ω s).to_real) μ := +begin + rw condexp_kernel, + exact integrable_to_real_cond_distrib (ae_measurable_id'' μ hm) hs, +end + +end integrability + +/-- The conditional expectation of `f` with respect to a σ-algebra `m` is almost everywhere equal to +the integral `∫ y, f y ∂(condexp_kernel μ m ω)`. -/ +lemma condexp_ae_eq_integral_condexp_kernel [normed_space ℝ F] [complete_space F] + (hm : m ≤ mΩ) (hf_int : integrable f μ) : + μ[f | m] =ᵐ[μ] λ ω, ∫ y, f y ∂(condexp_kernel μ m ω) := +begin + have hX : @measurable Ω Ω mΩ m id := measurable_id.mono le_rfl hm, + rw condexp_kernel, + refine eventually_eq.trans _ (condexp_ae_eq_integral_cond_distrib_id hX hf_int), + simp only [measurable_space.comap_id, id.def], +end + +end probability_theory diff --git a/src/probability/notation.lean b/src/probability/notation.lean index 23e11d9c09fa0..5e9d1e14b5f1c 100644 --- a/src/probability/notation.lean +++ b/src/probability/notation.lean @@ -15,6 +15,7 @@ measurable space `m0`, and another measurable space structure `m` with `hm : m - `𝔼[X|m]`: conditional expectation of `X` with respect to the measure `volume` and the measurable space `m`. The similar `P[X|m]` for a measure `P` is defined in measure_theory.function.conditional_expectation. +- `P⟦s|m⟧ = P[s.indicator (λ ω, (1 : ℝ)) | m]`, conditional probability of a set. - `X =ₐₛ Y`: `X =ᵐ[volume] Y` - `X ≤ₐₛ Y`: `X ≤ᵐ[volume] Y` - `∂P/∂Q = P.rn_deriv Q` @@ -26,6 +27,7 @@ We note that the notation `∂P/∂Q` applies to three different cases, namely, -/ open measure_theory +open_locale measure_theory -- We define notations `𝔼[f|m]` for the conditional expectation of `f` with respect to `m`. localized "notation (name := condexp.volume) `𝔼[` X `|` m `]` := @@ -36,6 +38,10 @@ localized "notation (name := condexp.probability) localized "notation (name := expected_value) `𝔼[` X `]` := ∫ a, X a" in probability_theory +localized "notation (name := condexp_indicator) + P `⟦` s `|` m `⟧` := measure_theory.condexp m P (s.indicator (λ ω, (1 : ℝ)))" + in probability_theory + localized "notation (name := eq_ae_volume) X ` =ₐₛ `:50 Y:50 := X =ᵐ[measure_theory.measure_space.volume] Y" in probability_theory From e137999b2c6f2be388f4cd3bbf8523de1910cd2b Mon Sep 17 00:00:00 2001 From: Moritz Doll Date: Wed, 31 May 2023 13:56:43 +0000 Subject: [PATCH 1122/1291] feat(analysis/schwartz_space): Multiplication of Schwartz function and functions of temperate growth (#18649) --- src/analysis/schwartz_space.lean | 181 +++++++++++++++++++++++++++++++ 1 file changed, 181 insertions(+) diff --git a/src/analysis/schwartz_space.lean b/src/analysis/schwartz_space.lean index b7971ba0335e9..085436610ee0f 100644 --- a/src/analysis/schwartz_space.lean +++ b/src/analysis/schwartz_space.lean @@ -59,6 +59,8 @@ Schwartz space, tempered distributions noncomputable theory +open_locale big_operators nat + variables {𝕜 𝕜' D E F G : Type*} variables [normed_add_comm_group E] [normed_space ℝ E] @@ -518,6 +520,37 @@ instance : topological_space.first_countable_topology (𝓢(E, F)) := end topology +section temperate_growth + +/-! ### Functions of temperate growth -/ + +/-- A function is called of temperate growth if it is smooth and all iterated derivatives are +polynomially bounded. -/ +def _root_.function.has_temperate_growth (f : E → F) : Prop := + cont_diff ℝ ⊤ f ∧ ∀ n : ℕ, ∃ (k : ℕ) (C : ℝ), ∀ x, ‖iterated_fderiv ℝ n f x‖ ≤ C * (1 + ‖x‖)^k + +lemma _root_.function.has_temperate_growth.norm_iterated_fderiv_le_uniform_aux {f : E → F} + (hf_temperate : f.has_temperate_growth) (n : ℕ) : + ∃ (k : ℕ) (C : ℝ) (hC : 0 ≤ C), ∀ (N : ℕ) (hN : N ≤ n) (x : E), + ‖iterated_fderiv ℝ N f x‖ ≤ C * (1 + ‖x‖)^k := +begin + choose k C f using hf_temperate.2, + use (finset.range (n+1)).sup k, + let C' := max (0 : ℝ) ((finset.range (n+1)).sup' (by simp) C), + have hC' : 0 ≤ C' := by simp only [le_refl, finset.le_sup'_iff, true_or, le_max_iff], + use [C', hC'], + intros N hN x, + rw ← finset.mem_range_succ_iff at hN, + refine le_trans (f N x) (mul_le_mul _ _ (by positivity) hC'), + { simp only [finset.le_sup'_iff, le_max_iff], + right, + exact ⟨N, hN, rfl.le⟩ }, + refine pow_le_pow (by simp only [le_add_iff_nonneg_right, norm_nonneg]) _, + exact finset.le_sup hN, +end + +end temperate_growth + section clm /-! ### Construction of continuous linear maps between Schwartz spaces -/ @@ -596,6 +629,154 @@ mk_clm (λ f x, f x m) end eval_clm +section multiplication + +variables [normed_add_comm_group D] [normed_space ℝ D] +variables [normed_add_comm_group G] [normed_space ℝ G] + +/-- The map `f ↦ (x ↦ B (f x) (g x))` as a continuous `𝕜`-linear map on Schwartz space, +where `B` is a continuous `𝕜`-linear map and `g` is a function of temperate growth. -/ +def bilin_left_clm (B : E →L[ℝ] F →L[ℝ] G) {g : D → F} (hg : g.has_temperate_growth) : + 𝓢(D, E) →L[ℝ] 𝓢(D, G) := + -- Todo (after port): generalize to `B : E →L[𝕜] F →L[𝕜] G` and `𝕜`-linear +mk_clm (λ f x, B (f x) (g x)) + (λ _ _ _, by simp only [map_add, add_left_inj, pi.add_apply, eq_self_iff_true, + continuous_linear_map.add_apply]) + (λ _ _ _, by simp only [pi.smul_apply, continuous_linear_map.coe_smul', + continuous_linear_map.map_smul, ring_hom.id_apply]) + (λ f, (B.is_bounded_bilinear_map.cont_diff.restrict_scalars ℝ).comp (f.smooth'.prod hg.1)) + (begin + -- Porting note: rewrite this proof with `rel_congr` + rintro ⟨k, n⟩, + rcases hg.norm_iterated_fderiv_le_uniform_aux n with ⟨l, C, hC, hgrowth⟩, + use [finset.Iic (l+k,n), ‖B‖ * (n + 1) * n.choose (n / 2) * (C * 2^(l + k)), by positivity], + intros f x, + have hxk : 0 ≤ ‖x‖^k := by positivity, + have hnorm_mul := + continuous_linear_map.norm_iterated_fderiv_le_of_bilinear B f.smooth' hg.1 x le_top, + refine le_trans (mul_le_mul_of_nonneg_left hnorm_mul hxk) _, + rw [← mul_assoc (‖x‖^k), mul_comm (‖x‖^k)], + simp_rw [mul_assoc (‖B‖)], + refine mul_le_mul_of_nonneg_left _ (by positivity), + rw [finset.mul_sum], + have : ∑ (x_1 : ℕ) in finset.range (n + 1), (1 : ℝ) = n + 1 := by simp, + repeat { rw [mul_assoc ((n : ℝ) + 1)] }, + rw [← this, finset.sum_mul], + refine finset.sum_le_sum (λ i hi, _), + simp only [one_mul], + rw [← mul_assoc, mul_comm (‖x‖^k), mul_assoc, mul_assoc, mul_assoc], + refine mul_le_mul _ _ (by positivity) (by positivity), + { norm_cast, + exact i.choose_le_middle n }, + specialize hgrowth (n - i) (by simp only [tsub_le_self]) x, + rw [← mul_assoc], + refine le_trans (mul_le_mul_of_nonneg_left hgrowth (by positivity)) _, + rw [mul_comm _ (C * _), mul_assoc, mul_assoc C], + refine mul_le_mul_of_nonneg_left _ hC, + nth_rewrite 1 mul_comm, + rw [← mul_assoc], + rw finset.mem_range_succ_iff at hi, + change i ≤ (l + k, n).snd at hi, + refine le_trans _ (one_add_le_sup_seminorm_apply le_rfl hi f x ), + refine mul_le_mul_of_nonneg_right _ (norm_nonneg _), + rw [pow_add], + refine mul_le_mul_of_nonneg_left _ (by positivity), + refine pow_le_pow_of_le_left (norm_nonneg _) _ _, + simp only [zero_le_one, le_add_iff_nonneg_left], + end) + +end multiplication + +section comp + +variables (𝕜) +variables [is_R_or_C 𝕜] +variables [normed_add_comm_group D] [normed_space ℝ D] +variables [normed_add_comm_group G] [normed_space ℝ G] +variables [normed_space 𝕜 F] [smul_comm_class ℝ 𝕜 F] +variables [normed_space 𝕜 G] [smul_comm_class ℝ 𝕜 G] + +/-- Composition with a function on the right is a continuous linear map on Schwartz space +provided that the function is temperate and growths polynomially near infinity. -/ +def comp_clm {g : D → E} (hg : g.has_temperate_growth) + (hg_upper : ∃ (k : ℕ) (C : ℝ), ∀ x, ‖x‖ ≤ C * (1 + ‖g x‖)^k ) : + 𝓢(E, F) →L[𝕜] 𝓢(D, F) := +mk_clm (λ f x, (f (g x))) + (λ _ _ _, by simp only [add_left_inj, pi.add_apply, eq_self_iff_true]) + (λ _ _ _, rfl) + (λ f, f.smooth'.comp hg.1) + (begin + rintros ⟨k, n⟩, + rcases hg.norm_iterated_fderiv_le_uniform_aux n with ⟨l, C, hC, hgrowth⟩, + rcases hg_upper with ⟨kg, Cg, hg_upper'⟩, + have hCg : 1 ≤ 1 + Cg := + begin + refine le_add_of_nonneg_right _, + specialize hg_upper' 0, + rw [norm_zero] at hg_upper', + refine nonneg_of_mul_nonneg_left hg_upper' (by positivity), + end, + let k' := kg * (k + l * n), + use [finset.Iic (k',n), (1 + Cg) ^ (k + l * n) * ((C + 1) ^ n * n! * 2 ^ k'), by positivity], + intros f x, + let seminorm_f := ((finset.Iic (k',n)).sup (schwartz_seminorm_family 𝕜 _ _)) f, + have hg_upper'' : (1 + ‖x‖)^(k + l * n) ≤ (1 + Cg)^(k + l*n) * (1 + ‖g x‖)^k' := + begin + rw [pow_mul, ← mul_pow], + refine pow_le_pow_of_le_left (by positivity) _ _, + rw [add_mul], + refine add_le_add _ (hg_upper' x), + nth_rewrite 0 ← one_mul (1 : ℝ), + refine mul_le_mul (le_refl _) (one_le_pow_of_one_le _ _) zero_le_one zero_le_one, + simp only [le_add_iff_nonneg_right, norm_nonneg], + end, + have hbound : ∀ i, i ≤ n → ‖iterated_fderiv ℝ i f (g x)‖ ≤ + 2 ^ k' * seminorm_f / ((1 + ‖g x‖) ^ k'):= + begin + intros i hi, + have hpos : 0 < (1 + ‖g x‖) ^ k' := by positivity, + rw le_div_iff' hpos, + change i ≤ (k', n).snd at hi, + exact one_add_le_sup_seminorm_apply le_rfl hi _ _, + end, + have hgrowth' : ∀ (N : ℕ) (hN₁ : 1 ≤ N) (hN₂ : N ≤ n), + ‖iterated_fderiv ℝ N g x‖ ≤ ((C + 1) * (1 + ‖x‖)^l)^N := + begin + intros N hN₁ hN₂, + refine (hgrowth N hN₂ x).trans _, + rw [mul_pow], + have hN₁' := (lt_of_lt_of_le zero_lt_one hN₁).ne.symm, + refine mul_le_mul _ _ (by positivity) (by positivity), + { exact le_trans (by simp [hC]) (le_self_pow (by simp [hC]) hN₁'), }, + { refine le_self_pow (one_le_pow_of_one_le _ l) hN₁', + simp only [le_add_iff_nonneg_right, norm_nonneg] }, + end, + have := norm_iterated_fderiv_comp_le f.smooth' hg.1 le_top x hbound hgrowth', + have hxk : ‖x‖^k ≤ (1 + ‖x‖)^k := + pow_le_pow_of_le_left (norm_nonneg _) (by simp only [zero_le_one, le_add_iff_nonneg_left]) _, + refine le_trans (mul_le_mul hxk this (by positivity) (by positivity)) _, + have rearrange : + (1 + ‖x‖) ^ k * (n! * (2 ^ k' * seminorm_f / (1 + ‖g x‖) ^ k') * + ((C + 1) * (1 + ‖x‖) ^ l) ^ n) = + ((1 + ‖x‖)^(k + l * n) / (1 + ‖g x‖) ^ k') * ((C + 1)^n * n! * 2^k' * seminorm_f) := + begin + rw [mul_pow, pow_add, ← pow_mul], + ring, + end, + rw rearrange, + have hgxk' : 0 < (1 + ‖g x‖) ^ k' := by positivity, + rw ← div_le_iff hgxk' at hg_upper'', + have hpos : 0 ≤ (C + 1) ^ n * n! * 2 ^ k' * seminorm_f := + begin + have : 0 ≤ seminorm_f := map_nonneg _ _, + positivity, + end, + refine le_trans (mul_le_mul_of_nonneg_right hg_upper'' hpos) _, + rw [← mul_assoc], + end) + +end comp + section derivatives /-! ### Derivatives of Schwartz functions -/ From cca40788df1b8755d5baf17ab2f27dacc2e17acb Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Wed, 31 May 2023 19:23:50 +0000 Subject: [PATCH 1123/1291] feat (number_theory/zeta_function): relate to Dirichlet series (#19131) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR adds two important properties of the Riemann zeta function: firstly, it is differentiable away from `s = 1` (which is easy for `s ≠ 0`, but much harder at `s = 0`); secondly, for `1 < re s` the zeta function is given by the usual Dirichlet series. Just for fun, we also add a formal statement of the Riemann hypothesis. --- .../special_functions/gamma/basic.lean | 13 ++ .../special_functions/gamma/beta.lean | 8 + src/number_theory/zeta_function.lean | 203 +++++++++++++++++- 3 files changed, 215 insertions(+), 9 deletions(-) diff --git a/src/analysis/special_functions/gamma/basic.lean b/src/analysis/special_functions/gamma/basic.lean index 693d12acb5cf7..07381ea31b077 100644 --- a/src/analysis/special_functions/gamma/basic.lean +++ b/src/analysis/special_functions/gamma/basic.lean @@ -447,6 +447,19 @@ end end Gamma_has_deriv +/-- At `s = 0`, the Gamma function has a simple pole with residue 1. -/ +lemma tendsto_self_mul_Gamma_nhds_zero : tendsto (λ z : ℂ, z * Gamma z) (𝓝[≠] 0) (𝓝 1) := +begin + rw (show 𝓝 (1 : ℂ) = 𝓝 (Gamma (0 + 1)), by simp only [zero_add, complex.Gamma_one]), + convert (tendsto.mono_left _ nhds_within_le_nhds).congr' + (eventually_eq_of_mem self_mem_nhds_within complex.Gamma_add_one), + refine continuous_at.comp _ (continuous_id.add continuous_const).continuous_at, + refine (complex.differentiable_at_Gamma _ (λ m, _)).continuous_at, + rw [zero_add, ←of_real_nat_cast, ←of_real_neg, ←of_real_one, ne.def, of_real_inj], + refine (lt_of_le_of_lt _ zero_lt_one).ne', + exact neg_nonpos.mpr (nat.cast_nonneg _), +end + end complex namespace real diff --git a/src/analysis/special_functions/gamma/beta.lean b/src/analysis/special_functions/gamma/beta.lean index a30c97499cc4f..6c24104b458b7 100644 --- a/src/analysis/special_functions/gamma/beta.lean +++ b/src/analysis/special_functions/gamma/beta.lean @@ -476,6 +476,14 @@ begin { rintro ⟨m, rfl⟩, exact Gamma_neg_nat_eq_zero m }, end +/-- A weaker, but easier-to-apply, version of `complex.Gamma_ne_zero`. -/ +lemma Gamma_ne_zero_of_re_pos {s : ℂ} (hs : 0 < re s) : Gamma s ≠ 0 := +begin + refine Gamma_ne_zero (λ m, _), + contrapose! hs, + simpa only [hs, neg_re, ←of_real_nat_cast, of_real_re, neg_nonpos] using nat.cast_nonneg _, +end + end complex namespace real diff --git a/src/number_theory/zeta_function.lean b/src/number_theory/zeta_function.lean index 20b3b7a2b3991..2a1c9d2101337 100644 --- a/src/number_theory/zeta_function.lean +++ b/src/number_theory/zeta_function.lean @@ -3,7 +3,7 @@ Copyright (c) 2023 David Loeffler. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: David Loeffler -/ -import analysis.mellin_transform +import analysis.special_functions.gamma.beta import number_theory.modular_forms.jacobi_theta /-! @@ -24,13 +24,16 @@ I haven't checked exactly what they are). ## Main results: * `differentiable_completed_zeta₀` : the function `Λ₀(s)` is entire. -* `differentiable_at_completed_zeta`: the function `Λ(s)` is differentiable away from `s = 0` and +* `differentiable_at_completed_zeta` : the function `Λ(s)` is differentiable away from `s = 0` and `s = 1`. +* `differentiable_at_riemann_zeta` : the function `ζ(s)` is differentiable away from `s = 1`. +* `zeta_eq_tsum_of_one_lt_re` : for `1 < re s`, we have + `ζ(s) = ∑' (n : ℕ), 1 / (n + 1) ^ s`. ## Outline of proofs: We define two related functions on the reals, `zeta_kernel₁` and `zeta_kernel₂`. The first is -`(Θ (t * I) - 1) / 2`, where `θ` is Jacobi's theta function; its Mellin transform is exactly the +`(θ (t * I) - 1) / 2`, where `θ` is Jacobi's theta function; its Mellin transform is exactly the completed zeta function. The second is obtained by subtracting a linear combination of powers on the interval `Ioc 0 1` to give a function with exponential decay at both `0` and `∞`. We then define `riemann_completed_zeta₀` as the Mellin transform of the second zeta kernel, and define @@ -67,9 +70,7 @@ details. -/ (λ s : ℂ, ↑π ^ (s / 2) * riemann_completed_zeta s / Gamma (s / 2)) 0 (-1 / 2) /- Note the next lemma is true by definition; what's hard is to show that with this definition, `ζ` -is continuous (and indeed analytic) at 0. - -TODO: Prove this -- this will be in a future PR. -/ +is continuous (and indeed analytic) at 0, see `differentiable_riemann_zeta` below. -/ /-- We have `ζ(0) = -1 / 2`. -/ lemma riemann_zeta_zero : riemann_zeta 0 = -1 / 2 := begin @@ -292,6 +293,94 @@ begin exact differentiable_at_id.sub (differentiable_at_const _) }, end +/-- The Riemann zeta function is differentiable away from `s = 1`. -/ +theorem differentiable_at_riemann_zeta {s : ℂ} (hs' : s ≠ 1) : + differentiable_at ℂ riemann_zeta s := +begin + /- First claim: the result holds at `t` for `t ≠ 0`. Note we will need to use this for the case + `s = 0` also, as a hypothesis for the removable-singularity criterion. -/ + have c1 : ∀ (t : ℂ) (ht : t ≠ 0) (ht' : t ≠ 1), differentiable_at ℂ + (λ u : ℂ, ↑π ^ (u / 2) * riemann_completed_zeta u / Gamma (u / 2)) t, + { intros t ht ht', + apply differentiable_at.mul, + { refine (differentiable_at.const_cpow _ _).mul (differentiable_at_completed_zeta ht ht'), + { exact differentiable_at.div_const differentiable_at_id _ }, + { exact or.inl (of_real_ne_zero.mpr pi_pos.ne') } }, + { refine differentiable_one_div_Gamma.differentiable_at.comp t _, + exact differentiable_at.div_const differentiable_at_id _ } }, + /- Second claim: the limit at `s = 0` exists and is equal to `-1 / 2`. -/ + have c2 : tendsto (λ s : ℂ, ↑π ^ (s / 2) * riemann_completed_zeta s / Gamma (s / 2)) + (𝓝[≠] 0) (𝓝 $ -1 / 2), + { have h1 : tendsto (λ z : ℂ, (π : ℂ) ^ (z / 2)) (𝓝 0) (𝓝 1), + { convert (continuous_at_const_cpow (of_real_ne_zero.mpr pi_pos.ne')).comp _, + { simp_rw [function.comp_app, zero_div, cpow_zero] }, + { exact continuous_at_id.div continuous_at_const two_ne_zero } }, + suffices h2 : tendsto (λ z, riemann_completed_zeta z / Gamma (z / 2)) (𝓝[≠] 0) (𝓝 $ -1 / 2), + { convert (h1.mono_left nhds_within_le_nhds).mul h2, + { ext1 x, rw mul_div }, { simp only [one_mul] } }, + suffices h3 : tendsto (λ z, (riemann_completed_zeta z * (z / 2)) / (z / 2 * Gamma (z / 2))) + (𝓝[≠] 0) (𝓝 $ -1 / 2), + { refine tendsto.congr' (eventually_eq_of_mem self_mem_nhds_within (λ z hz, _)) h3, + rw [←div_div, mul_div_cancel _ (div_ne_zero hz two_ne_zero)] }, + have h4 : tendsto (λ z : ℂ, z / 2 * Gamma (z / 2)) (𝓝[≠] 0) (𝓝 1), + { refine tendsto_self_mul_Gamma_nhds_zero.comp _, + rw [tendsto_nhds_within_iff, (by simp : 𝓝 (0 : ℂ) = 𝓝 (0 / 2))], + exact ⟨(tendsto_id.div_const _).mono_left nhds_within_le_nhds, + eventually_of_mem self_mem_nhds_within (λ x hx, div_ne_zero hx two_ne_zero)⟩ }, + suffices : tendsto (λ z, riemann_completed_zeta z * z / 2) (𝓝[≠] 0) (𝓝 (-1 / 2 : ℂ)), + { have := this.div h4 one_ne_zero, + simp_rw [div_one, mul_div_assoc] at this, + exact this }, + refine tendsto.div _ tendsto_const_nhds two_ne_zero, + simp_rw [riemann_completed_zeta, add_mul, sub_mul], + rw show 𝓝 (-1 : ℂ) = 𝓝 (0 - 1 + 0), by rw [zero_sub, add_zero], + refine (tendsto.sub _ _).add _, + { refine tendsto.mono_left _ nhds_within_le_nhds, + have : continuous_at riemann_completed_zeta₀ 0, + from (differentiable_completed_zeta₀).continuous.continuous_at, + simpa only [id.def, mul_zero] using tendsto.mul this tendsto_id }, + { refine tendsto_const_nhds.congr' (eventually_eq_of_mem self_mem_nhds_within (λ t ht, _)), + simp_rw one_div_mul_cancel ht }, + { refine tendsto.mono_left _ nhds_within_le_nhds, + suffices : continuous_at (λ z : ℂ, 1 / (z - 1)) 0, + by simpa only [id.def, mul_zero] using tendsto.mul this tendsto_id, + refine continuous_at_const.div (continuous_at_id.sub continuous_at_const) _, + simpa only [zero_sub] using neg_ne_zero.mpr one_ne_zero } }, + -- Now the main proof. + rcases ne_or_eq s 0 with hs | rfl, + { -- The easy case: `s ≠ 0` + have : {(0 : ℂ)}ᶜ ∈ 𝓝 s, from is_open_compl_singleton.mem_nhds hs, + refine (c1 s hs hs').congr_of_eventually_eq (eventually_eq_of_mem this (λ x hx, _)), + unfold riemann_zeta, + apply function.update_noteq hx }, + { -- The hard case: `s = 0`. + rw [riemann_zeta, ←(lim_eq_iff ⟨-1 / 2, c2⟩).mpr c2], + have S_nhds : {(1 : ℂ)}ᶜ ∈ 𝓝 (0 : ℂ), from is_open_compl_singleton.mem_nhds hs', + refine ((complex.differentiable_on_update_lim_of_is_o S_nhds + (λ t ht, (c1 t ht.2 ht.1).differentiable_within_at) _) 0 hs').differentiable_at S_nhds, + simp only [zero_div, div_zero, complex.Gamma_zero, mul_zero, cpow_zero, sub_zero], + -- Remains to show completed zeta is `o (s ^ (-1))` near 0. + refine (is_O_const_of_tendsto c2 $ one_ne_zero' ℂ).trans_is_o _, + rw is_o_iff_tendsto', + { exact tendsto.congr (λ x, by rw [←one_div, one_div_one_div]) nhds_within_le_nhds }, + { exact eventually_of_mem self_mem_nhds_within (λ x hx hx', (hx $ inv_eq_zero.mp hx').elim) } } +end + +/-- The trivial zeroes of the zeta function. -/ +lemma riemann_zeta_neg_two_mul_nat_add_one (n : ℕ) : riemann_zeta (-2 * (n + 1)) = 0 := +begin + have : (-2 : ℂ) * (n + 1) ≠ 0, + from mul_ne_zero (neg_ne_zero.mpr two_ne_zero) (nat.cast_add_one_ne_zero n), + rw [riemann_zeta, function.update_noteq this, + (show (-2) * ((n : ℂ) + 1) / 2 = -↑(n + 1), by { push_cast, ring }), + complex.Gamma_neg_nat_eq_zero, div_zero], +end + +/-- A formal statement of the Riemann hypothesis – constructing a term of this type is worth a +million dollars. -/ +def riemann_hypothesis : Prop := +∀ (s : ℂ) (hs : riemann_completed_zeta s = 0) (hs' : ¬∃ (n : ℕ), s = -2 * (n + 1)), s.re = 1 / 2 + /-! ## Relating the Mellin transforms of the two zeta kernels -/ @@ -310,7 +399,7 @@ end /-- Evaluate the Mellin transform of the "fudge factor" in `zeta_kernel₂` -/ lemma has_mellin_one_div_sqrt_sub_one_div_two_Ioc {s : ℂ} (hs : 1 / 2 < s.re) : - has_mellin ((Ioc 0 1).indicator (λ t, (1 - 1 / (↑(sqrt t) : ℂ)) / 2)) s + has_mellin ((Ioc 0 1).indicator (λ t, (1 - 1 / (sqrt t : ℂ)) / 2)) s (1 / (2 * s) - 1 / (2 * s - 1)) := begin have step1 : has_mellin (indicator (Ioc 0 1) (λ t, 1 - 1 / ↑(sqrt t) : ℝ → ℂ)) s @@ -319,8 +408,8 @@ begin have b := has_mellin_one_div_sqrt_Ioc hs, simpa only [a.2, b.2, ←indicator_sub] using has_mellin_sub a.1 b.1 }, -- todo: implement something like "indicator.const_div" (blocked by the port for now) - rw (show (Ioc 0 1).indicator (λ t, (1 - 1 / (↑(sqrt t) : ℂ)) / 2) = - λ t, ((Ioc 0 1).indicator (λ t, (1 - 1 / (↑(sqrt t) : ℂ))) t) / 2, + rw (show (Ioc 0 1).indicator (λ t, (1 - 1 / (sqrt t : ℂ)) / 2) = + λ t, ((Ioc 0 1).indicator (λ t, (1 - 1 / (sqrt t : ℂ))) t) / 2, by { ext1 t, simp_rw [div_eq_inv_mul, indicator_mul_right] }), simp_rw [has_mellin, mellin_div_const, step1.2, sub_div, div_div], refine ⟨step1.1.div_const _, _⟩, @@ -349,3 +438,99 @@ begin rw mul_div_cancel' _ (two_ne_zero' ℂ), abel end + +/-! +## Relating the first zeta kernel to the Dirichlet series +-/ + +/-- Auxiliary lemma for `mellin_zeta_kernel₁_eq_tsum`, computing the Mellin transform of an +individual term in the series. -/ +lemma integral_cpow_mul_exp_neg_pi_mul_sq {s : ℂ} (hs : 0 < s.re) (n : ℕ) : + ∫ t : ℝ in Ioi 0, (t : ℂ) ^ (s - 1) * rexp (-π * t * (n + 1) ^ 2) = + ↑π ^ -s * complex.Gamma s * (1 / (n + 1) ^ (2 * s)) := +begin + rw [complex.Gamma_eq_integral hs, Gamma_integral_eq_mellin], + conv_rhs { congr, rw [←smul_eq_mul, ←mellin_comp_mul_left _ _ pi_pos] }, + have : (1 / ((n : ℂ) + 1) ^ (2 * s)) = ↑(((n : ℝ) + 1) ^ (2 : ℝ)) ^ (-s), + { rw [(by push_cast: ((n : ℂ) + 1) = ↑( (n : ℝ) + 1)), + (by push_cast : (2 * s) = (↑(2 : ℝ) * s)), + cpow_mul_of_real_nonneg, one_div, cpow_neg], + rw [←nat.cast_succ], + exact nat.cast_nonneg _ }, + conv_rhs { rw [this, mul_comm, ←smul_eq_mul] }, + rw [← mellin_comp_mul_right _ _ (show 0 < ((n : ℝ) + 1) ^ (2 : ℝ), by positivity)], + refine set_integral_congr measurable_set_Ioi (λ t ht, _), + simp_rw smul_eq_mul, + congr' 3, + conv_rhs { rw [←nat.cast_two, rpow_nat_cast] }, + ring +end + +lemma mellin_zeta_kernel₁_eq_tsum {s : ℂ} (hs : 1 / 2 < s.re): + mellin zeta_kernel₁ s = π ^ (-s) * Gamma s * ∑' (n : ℕ), 1 / (n + 1) ^ (2 * s) := +begin + let bd : ℕ → ℝ → ℝ := λ n t, t ^ (s.re - 1) * exp (-π * t * (n + 1) ^ 2), + let f : ℕ → ℝ → ℂ := λ n t, t ^ (s - 1) * exp (-π * t * (n + 1) ^ 2), + have hm : measurable_set (Ioi (0:ℝ)), from measurable_set_Ioi, + have h_norm : ∀ (n : ℕ) {t : ℝ} (ht : 0 < t), ‖f n t‖ = bd n t, + { intros n t ht, + rw [norm_mul, complex.norm_eq_abs, complex.norm_eq_abs, complex.abs_of_nonneg (exp_pos _).le, + abs_cpow_eq_rpow_re_of_pos ht, sub_re, one_re] }, + have hf_meas : ∀ (n : ℕ), ae_strongly_measurable (f n) (volume.restrict $ Ioi 0), + { intro n, + refine (continuous_on.mul _ _).ae_strongly_measurable hm, + { exact (continuous_at.continuous_on + (λ x hx, continuous_at_of_real_cpow_const _ _ $ or.inr $ ne_of_gt hx)) }, + { apply continuous.continuous_on, + exact continuous_of_real.comp (continuous_exp.comp + ((continuous_const.mul continuous_id').mul continuous_const)) } }, + have h_le : ∀ (n : ℕ), ∀ᵐ (t : ℝ) ∂volume.restrict (Ioi 0), ‖f n t‖ ≤ bd n t, + from λ n, (ae_restrict_iff' hm).mpr (ae_of_all _ (λ t ht, le_of_eq (h_norm n ht))), + have h_sum0 : ∀ {t : ℝ} (ht : 0 < t), has_sum (λ n, f n t) (t ^ (s - 1) * zeta_kernel₁ t), + { intros t ht, + have := (has_sum_of_real.mpr (summable_exp_neg_pi_mul_nat_sq ht).has_sum).mul_left + ((t : ℂ) ^ (s - 1)), + simpa only [of_real_mul, ←mul_assoc, of_real_bit0, of_real_one, mul_comm _ (2 : ℂ), + of_real_sub, of_real_one, of_real_tsum] using this }, + have h_sum' : ∀ᵐ (t : ℝ) ∂volume.restrict (Ioi 0), has_sum (λ (n : ℕ), f n t) + (t ^ (s - 1) * zeta_kernel₁ t), + from (ae_restrict_iff' hm).mpr (ae_of_all _ (λ t ht, h_sum0 ht)), + have h_sum : ∀ᵐ (t : ℝ) ∂volume.restrict (Ioi 0), summable (λ n : ℕ, bd n t), + { refine (ae_restrict_iff' hm).mpr (ae_of_all _ (λ t ht, _)), + simpa only [λ n, h_norm n ht] using summable_norm_iff.mpr (h_sum0 ht).summable }, + have h_int : integrable (λ t : ℝ, ∑' (n : ℕ), bd n t) (volume.restrict (Ioi 0)), + { refine integrable_on.congr_fun (mellin_convergent_of_is_O_rpow_exp pi_pos + locally_integrable_zeta_kernel₁ is_O_at_top_zeta_kernel₁ is_O_zero_zeta_kernel₁ hs).norm + (λ t ht, _) hm, + simp_rw [tsum_mul_left, norm_smul, complex.norm_eq_abs ((t : ℂ) ^ _), + abs_cpow_eq_rpow_re_of_pos ht, sub_re, one_re], + rw [zeta_kernel₁, ←of_real_tsum, complex.norm_eq_abs, complex.abs_of_nonneg], + exact tsum_nonneg (λ n, (exp_pos _).le) }, + simpa only [integral_cpow_mul_exp_neg_pi_mul_sq (one_half_pos.trans hs), tsum_mul_left] using + (has_sum_integral_of_dominated_convergence bd hf_meas h_le h_sum h_int h_sum').tsum_eq.symm, +end + +lemma completed_zeta_eq_tsum_of_one_lt_re {s : ℂ} (hs : 1 < re s) : + riemann_completed_zeta s = π ^ (-s / 2) * Gamma (s / 2) * ∑' (n : ℕ), 1 / (n + 1) ^ s := +begin + rw [completed_zeta_eq_mellin_of_one_lt_re hs, mellin_zeta_kernel₁_eq_tsum, neg_div, + mul_div_cancel' _ (two_ne_zero' ℂ)], + rw (show s / 2 = ↑(2⁻¹ : ℝ) * s, by { push_cast, rw mul_comm, refl }), + rwa [of_real_mul_re, ←div_eq_inv_mul, div_lt_div_right (zero_lt_two' ℝ)] +end + +/-- The Riemann zeta function agrees with the naive Dirichlet-series definition when the latter +converges. (Note that this is false without the assumption: when `re s ≤ 1` the sum is divergent, +and we use a different definition to obtain the analytic continuation to all `s`.) -/ +theorem zeta_eq_tsum_of_one_lt_re {s : ℂ} (hs : 1 < re s) : + riemann_zeta s = ∑' (n : ℕ), 1 / (n + 1) ^ s := +begin + have : s ≠ 0, by { contrapose! hs, rw [hs, zero_re], exact zero_le_one }, + rw [riemann_zeta, function.update_noteq this, completed_zeta_eq_tsum_of_one_lt_re hs, + ←mul_assoc, neg_div, cpow_neg, mul_inv_cancel_left₀, mul_div_cancel_left], + { apply Gamma_ne_zero_of_re_pos, + rw [←of_real_one, ←of_real_bit0, div_eq_mul_inv, ←of_real_inv, mul_comm, of_real_mul_re], + exact mul_pos (inv_pos_of_pos two_pos) (zero_lt_one.trans hs), }, + { rw [ne.def, cpow_eq_zero_iff, not_and_distrib, ←ne.def, of_real_ne_zero], + exact or.inl (pi_pos.ne') } +end From 04e80bb7e8510958cd9aacd32fe2dc147af0b9f1 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Thu, 1 Jun 2023 05:11:40 +0000 Subject: [PATCH 1124/1291] refactor(number_theory/liouville/liouville_number): review API, golf (#19126) * Protect `liouville.irrational` and `liouville.transcendental`. * Sync file name with definition name (Liouville constant vs Liouville number). * Move auxiliary definitions and lemmas about Liouville number to `liouville_number` namespace. * Rename auxiliary definitions, golf proofs (where it doesn't affect readability). --- docs/overview.yaml | 2 +- src/number_theory/liouville/basic.lean | 4 +- ...le_constant.lean => liouville_number.lean} | 139 +++++++++--------- 3 files changed, 73 insertions(+), 72 deletions(-) rename src/number_theory/liouville/{liouville_constant.lean => liouville_number.lean} (53%) diff --git a/docs/overview.yaml b/docs/overview.yaml index f5178f64aff7e..d1653f00c6261 100644 --- a/docs/overview.yaml +++ b/docs/overview.yaml @@ -137,7 +137,7 @@ General algebra: ring of Witt vectors: 'witt_vector' perfection of a ring: 'ring.perfection' Transcendental numbers: - Liouville's theorem on existence of transcendental numbers: liouville.is_transcendental + Liouville's theorem on existence of transcendental numbers: 'transcendental_liouville_number' Representation theory: representation : 'representation' diff --git a/src/number_theory/liouville/basic.lean b/src/number_theory/liouville/basic.lean index 3fd7bc77a56c6..438cb18b980fb 100644 --- a/src/number_theory/liouville/basic.lean +++ b/src/number_theory/liouville/basic.lean @@ -29,7 +29,7 @@ def liouville (x : ℝ) := ∀ n : ℕ, ∃ a b : ℤ, 1 < b ∧ x ≠ a / b ∧ namespace liouville -@[protected] lemma irrational {x : ℝ} (h : liouville x) : irrational x := +protected lemma irrational {x : ℝ} (h : liouville x) : irrational x := begin -- By contradiction, `x = a / b`, with `a ∈ ℤ`, `0 < b ∈ ℕ` is a Liouville number, rintros ⟨⟨a, b, bN0, cop⟩, rfl⟩, @@ -169,7 +169,7 @@ begin end /-- **Liouville's Theorem** -/ -theorem transcendental {x : ℝ} (lx : liouville x) : +protected theorem transcendental {x : ℝ} (lx : liouville x) : transcendental ℤ x := begin -- Proceed by contradiction: if `x` is algebraic, then `x` is the root (`ef0`) of a diff --git a/src/number_theory/liouville/liouville_constant.lean b/src/number_theory/liouville/liouville_number.lean similarity index 53% rename from src/number_theory/liouville/liouville_constant.lean rename to src/number_theory/liouville/liouville_number.lean index 6cebf03edd18f..78d446cf58326 100644 --- a/src/number_theory/liouville/liouville_constant.lean +++ b/src/number_theory/liouville/liouville_number.lean @@ -10,7 +10,7 @@ import number_theory.liouville.basic This file contains a construction of a family of Liouville numbers, indexed by a natural number $m$. The most important property is that they are examples of transcendental real numbers. -This fact is recorded in `liouville.is_transcendental`. +This fact is recorded in `transcendental_liouville_number`. More precisely, for a real number $m$, Liouville's constant is $$ @@ -33,8 +33,6 @@ noncomputable theory open_locale nat big_operators open real finset -namespace liouville - /-- For a real number `m`, Liouville's constant is $$ @@ -45,78 +43,80 @@ if the series does not converge, then the sum of the series is defined to be zer -/ def liouville_number (m : ℝ) : ℝ := ∑' (i : ℕ), 1 / m ^ i! +namespace liouville_number + /-- -`liouville_number_initial_terms` is the sum of the first `k + 1` terms of Liouville's constant, +`liouville_number.partial_sum` is the sum of the first `k + 1` terms of Liouville's constant, i.e. $$ \sum_{i=0}^k\frac{1}{m^{i!}}. $$ -/ -def liouville_number_initial_terms (m : ℝ) (k : ℕ) : ℝ := ∑ i in range (k+1), 1 / m ^ i! +def partial_sum (m : ℝ) (k : ℕ) : ℝ := ∑ i in range (k+1), 1 / m ^ i! /-- -`liouville_number_tail` is the sum of the series of the terms in `liouville_number m` +`liouville_number.remainder` is the sum of the series of the terms in `liouville_number m` starting from `k+1`, i.e $$ \sum_{i=k+1}^\infty\frac{1}{m^{i!}}. $$ -/ -def liouville_number_tail (m : ℝ) (k : ℕ) : ℝ := ∑' i, 1 / m ^ (i + (k+1))! - -lemma liouville_number_tail_pos {m : ℝ} (hm : 1 < m) (k : ℕ) : - 0 < liouville_number_tail m k := --- replace `0` with the constantly zero series `∑ i : ℕ, 0` -calc (0 : ℝ) = ∑' i : ℕ, 0 : tsum_zero.symm - ... < liouville_number_tail m k : - -- to show that a series with non-negative terms has strictly positive sum it suffices - -- to prove that - tsum_lt_tsum_of_nonneg - -- 1. the terms of the zero series are indeed non-negative - (λ _, rfl.le) - -- 2. the terms of our series are non-negative - (λ i, one_div_nonneg.mpr (pow_nonneg (zero_le_one.trans hm.le) _)) - -- 3. one term of our series is strictly positive -- they all are, we use the first term - (one_div_pos.mpr (pow_pos (zero_lt_one.trans hm) (0 + (k + 1))!)) $ - -- 4. our series converges -- it does since it is the tail of a converging series, though - -- this is not the argument here. - summable_one_div_pow_of_le hm (λ i, trans le_self_add (nat.self_le_factorial _)) +def remainder (m : ℝ) (k : ℕ) : ℝ := ∑' i, 1 / m ^ (i + (k+1))! + +/-! +We start with simple observations. +-/ + +protected lemma summable {m : ℝ} (hm : 1 < m) : summable (λ i : ℕ, 1 / m ^ i!) := +summable_one_div_pow_of_le hm nat.self_le_factorial + +lemma remainder_summable {m : ℝ} (hm : 1 < m) (k : ℕ) : + summable (λ i : ℕ, 1 / m ^ (i + (k + 1))!) := +by convert (summable_nat_add_iff (k + 1)).2 (liouville_number.summable hm) + +lemma remainder_pos {m : ℝ} (hm : 1 < m) (k : ℕ) : 0 < remainder m k := +tsum_pos (remainder_summable hm k) (λ _, by positivity) 0 (by positivity) + +lemma partial_sum_succ (m : ℝ) (n : ℕ) : + partial_sum m (n + 1) = partial_sum m n + 1 / m ^ (n + 1)! := +sum_range_succ _ _ /-- Split the sum definining a Liouville number into the first `k` term and the rest. -/ -lemma liouville_number_eq_initial_terms_add_tail {m : ℝ} (hm : 1 < m) (k : ℕ) : - liouville_number m = liouville_number_initial_terms m k + - liouville_number_tail m k := -(sum_add_tsum_nat_add _ (summable_one_div_pow_of_le hm (λ i, i.self_le_factorial))).symm +lemma partial_sum_add_remainder {m : ℝ} (hm : 1 < m) (k : ℕ) : + partial_sum m k + remainder m k = liouville_number m := +sum_add_tsum_nat_add _ (liouville_number.summable hm) /-! We now prove two useful inequalities, before collecting everything together. -/ -/-- Partial inequality, works with `m ∈ ℝ` satisfying `1 < m`. -/ -lemma tsum_one_div_pow_factorial_lt (n : ℕ) {m : ℝ} (m1 : 1 < m) : - ∑' (i : ℕ), 1 / m ^ (i + (n + 1))! < (1 - 1 / m)⁻¹ * (1 / m ^ (n + 1)!) := +/-- An upper estimate on the remainder. This estimate works with `m ∈ ℝ` satisfying `1 < m` and is +stronger than the estimate `liouville_number.remainder_lt` below. However, the latter estimate is +more useful for the proof. -/ +lemma remainder_lt' (n : ℕ) {m : ℝ} (m1 : 1 < m) : + remainder m n < (1 - 1 / m)⁻¹ * (1 / m ^ (n + 1)!) := -- two useful inequalities -have m0 : 0 < m := (zero_lt_one.trans m1), -have mi : |1 / m| < 1 := - (le_of_eq (abs_of_pos (one_div_pos.mpr m0))).trans_lt ((div_lt_one m0).mpr m1), +have m0 : 0 < m := zero_lt_one.trans m1, +have mi : 1 / m < 1 := (div_lt_one m0).mpr m1, calc (∑' i, 1 / m ^ (i + (n + 1))!) < ∑' i, 1 / m ^ (i + (n + 1)!) : -- to show the strict inequality between these series, we prove that: - tsum_lt_tsum_of_nonneg - -- 1. the first series has non-negative terms - (λ b, one_div_nonneg.mpr (pow_nonneg m0.le _)) - -- 2. the second series dominates the first + tsum_lt_tsum + -- 1. the second series dominates the first (λ b, one_div_pow_le_one_div_pow_of_le m1.le (b.add_factorial_succ_le_factorial_add_succ n)) - -- 3. the term with index `i = 2` of the first series is strictly smaller than + -- 2. the term with index `i = 2` of the first series is strictly smaller than -- the corresponding term of the second series - (one_div_pow_strict_anti m1 (n.add_factorial_succ_lt_factorial_add_succ rfl.le)) + (one_div_pow_strict_anti m1 (n.add_factorial_succ_lt_factorial_add_succ le_rfl)) + -- 3. the first series is summable + (remainder_summable m1 n) -- 4. the second series is summable, since its terms grow quickly - (summable_one_div_pow_of_le m1 (λ j, nat.le.intro rfl)) -... = ∑' i, (1 / m) ^ i * (1 / m ^ (n + 1)!) : + (summable_one_div_pow_of_le m1 (λ j, le_self_add)) +... = ∑' i : ℕ, (1 / m) ^ i * (1 / m ^ (n + 1)!) : -- split the sum in the exponent and massage - by { congr, ext i, rw [pow_add, ← div_div, div_eq_mul_one_div, one_div_pow] } + by simp only [pow_add, one_div, mul_inv, inv_pow] -- factor the constant `(1 / m ^ (n + 1)!)` out of the series ... = (∑' i, (1 / m) ^ i) * (1 / m ^ (n + 1)!) : tsum_mul_right ... = (1 - 1 / m)⁻¹ * (1 / m ^ (n + 1)!) : -- the series if the geometric series - mul_eq_mul_right_iff.mpr (or.inl (tsum_geometric_of_abs_lt_1 mi)) + by rw [tsum_geometric_of_lt_1 (by positivity) mi] lemma aux_calc (n : ℕ) {m : ℝ} (hm : 2 ≤ m) : (1 - 1 / m)⁻¹ * (1 / m ^ (n + 1)!) ≤ 1 / (m ^ n!) ^ n := @@ -142,29 +142,36 @@ calc (1 - 1 / m)⁻¹ * (1 / m ^ (n + 1)!) ≤ 2 * (1 / m ^ (n + 1)!) : end ... = 1 / (m ^ n!) ^ n : congr_arg ((/) 1) (pow_mul m n! n) +/-- An upper estimate on the remainder. This estimate works with `m ∈ ℝ` satisfying `2 ≤ m` and is +weaker than the estimate `liouville_number.remainder_lt'` above. However, this estimate is +more useful for the proof. -/ +lemma remainder_lt (n : ℕ) {m : ℝ} (m2 : 2 ≤ m) : + remainder m n < 1 / (m ^ n!) ^ n := +(remainder_lt' n $ one_lt_two.trans_le m2).trans_le (aux_calc _ m2) + /-! Starting from here, we specialize to the case in which `m` is a natural number. -/ /-- The sum of the `k` initial terms of the Liouville number to base `m` is a ratio of natural numbers where the denominator is `m ^ k!`. -/ -lemma liouville_number_rat_initial_terms {m : ℕ} (hm : 0 < m) (k : ℕ) : -∃ p : ℕ, liouville_number_initial_terms m k = p / m ^ k! := +lemma partial_sum_eq_rat {m : ℕ} (hm : 0 < m) (k : ℕ) : + ∃ p : ℕ, partial_sum m k = p / m ^ k! := begin induction k with k h, - { exact ⟨1, by rw [liouville_number_initial_terms, range_one, sum_singleton, nat.cast_one]⟩ }, + { exact ⟨1, by rw [partial_sum, range_one, sum_singleton, nat.cast_one]⟩ }, { rcases h with ⟨p_k, h_k⟩, use p_k * (m ^ ((k + 1)! - k!)) + 1, - unfold liouville_number_initial_terms at h_k ⊢, - rw [sum_range_succ, h_k, div_add_div, div_eq_div_iff, add_mul], + rw [partial_sum_succ, h_k, div_add_div, div_eq_div_iff, add_mul], { norm_cast, - rw [add_mul, one_mul, nat.factorial_succ, - show k.succ * k! - k! = (k.succ - 1) * k!, by rw [tsub_mul, one_mul], - nat.succ_sub_one, add_mul, one_mul, pow_add], + rw [add_mul, one_mul, nat.factorial_succ, add_mul, one_mul, add_tsub_cancel_right, pow_add], simp [mul_assoc] }, - refine mul_ne_zero_iff.mpr ⟨_, _⟩, - all_goals { exact pow_ne_zero _ (nat.cast_ne_zero.mpr hm.ne.symm) } } + all_goals { positivity } } end -theorem is_liouville {m : ℕ} (hm : 2 ≤ m) : +end liouville_number + +open liouville_number + +theorem liouville_liouville_number {m : ℕ} (hm : 2 ≤ m) : liouville (liouville_number m) := begin -- two useful inequalities @@ -172,21 +179,15 @@ begin have m1 : 1 < (m : ℝ), { norm_cast, exact one_lt_two.trans_le hm }, intro n, -- the first `n` terms sum to `p / m ^ k!` - rcases liouville_number_rat_initial_terms (zero_lt_two.trans_le hm) n with ⟨p, hp⟩, + rcases partial_sum_eq_rat (zero_lt_two.trans_le hm) n with ⟨p, hp⟩, refine ⟨p, m ^ n!, one_lt_pow mZ1 n.factorial_ne_zero, _⟩, push_cast, -- separate out the sum of the first `n` terms and the rest - rw [liouville_number_eq_initial_terms_add_tail m1 n, - ← hp, add_sub_cancel', abs_of_nonneg (liouville_number_tail_pos m1 _).le], - exact ⟨((lt_add_iff_pos_right _).mpr (liouville_number_tail_pos m1 n)).ne.symm, - (tsum_one_div_pow_factorial_lt n m1).trans_le - (aux_calc _ (nat.cast_two.symm.le.trans (nat.cast_le.mpr hm)))⟩ + rw [← partial_sum_add_remainder m1 n, ← hp], + have hpos := remainder_pos m1 n, + simpa [abs_of_pos hpos, hpos.ne'] using @remainder_lt n m (by assumption_mod_cast) end -/- Placing this lemma outside of the `open/closed liouville`-namespace would allow to remove -`_root_.`, at the cost of some other small weirdness. -/ -lemma is_transcendental {m : ℕ} (hm : 2 ≤ m) : - _root_.transcendental ℤ (liouville_number m) := -transcendental (is_liouville hm) - -end liouville +lemma transcendental_liouville_number {m : ℕ} (hm : 2 ≤ m) : + transcendental ℤ (liouville_number m) := +(liouville_liouville_number hm).transcendental From 599fffe78f0e11eb6a034e834ec51882167b9688 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Thu, 1 Jun 2023 06:32:48 +0000 Subject: [PATCH 1125/1291] chore(*): add mathlib4 synchronization comments (#19135) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `analysis.calculus.deriv.zpow` * `analysis.quaternion` * `analysis.specific_limits.floor_pow` * `category_theory.sites.closed` * `group_theory.nilpotent` * `measure_theory.function.l1_space` * `measure_theory.group.prod` * `measure_theory.integral.integrable_on` * `measure_theory.measure.haar.basic` * `topology.homotopy.homotopy_group` * `topology.sheaves.local_predicate` --- src/analysis/calculus/deriv/zpow.lean | 3 +++ src/analysis/quaternion.lean | 3 +++ src/analysis/specific_limits/floor_pow.lean | 3 +++ src/category_theory/sites/closed.lean | 3 +++ src/group_theory/nilpotent.lean | 3 +++ src/measure_theory/function/l1_space.lean | 3 +++ src/measure_theory/group/prod.lean | 3 +++ src/measure_theory/integral/integrable_on.lean | 3 +++ src/measure_theory/measure/haar/basic.lean | 3 +++ src/topology/homotopy/homotopy_group.lean | 3 +++ src/topology/sheaves/local_predicate.lean | 3 +++ 11 files changed, 33 insertions(+) diff --git a/src/analysis/calculus/deriv/zpow.lean b/src/analysis/calculus/deriv/zpow.lean index 867c779eb3b23..507b92f86380a 100644 --- a/src/analysis/calculus/deriv/zpow.lean +++ b/src/analysis/calculus/deriv/zpow.lean @@ -9,6 +9,9 @@ import analysis.calculus.deriv.inv /-! # Derivatives of `x ^ m`, `m : ℤ` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove theorems about (iterated) derivatives of `x ^ m`, `m : ℤ`. For a more detailed overview of one-dimensional derivatives in mathlib, see the module docstring of diff --git a/src/analysis/quaternion.lean b/src/analysis/quaternion.lean index f00c116fe91f6..4fa0224f1e472 100644 --- a/src/analysis/quaternion.lean +++ b/src/analysis/quaternion.lean @@ -11,6 +11,9 @@ import topology.algebra.algebra /-! # Quaternions as a normed algebra +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define the following structures on the space `ℍ := ℍ[ℝ]` of quaternions: * inner product space; diff --git a/src/analysis/specific_limits/floor_pow.lean b/src/analysis/specific_limits/floor_pow.lean index 10db01964533b..1c7c1e209188a 100644 --- a/src/analysis/specific_limits/floor_pow.lean +++ b/src/analysis/specific_limits/floor_pow.lean @@ -10,6 +10,9 @@ import analysis.special_functions.pow.real /-! # Results on discretized exponentials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We state several auxiliary results pertaining to sequences of the form `⌊c^n⌋₊`. * `tendsto_div_of_monotone_of_tendsto_div_floor_pow`: If a monotone sequence `u` is such that diff --git a/src/category_theory/sites/closed.lean b/src/category_theory/sites/closed.lean index 2a3ef6114eee2..a2cf4f46df1f3 100644 --- a/src/category_theory/sites/closed.lean +++ b/src/category_theory/sites/closed.lean @@ -9,6 +9,9 @@ import order.closure /-! # Closed sieves +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A natural closure operator on sieves is a closure operator on `sieve X` for each `X` which commutes with pullback. We show that a Grothendieck topology `J` induces a natural closure operator, and define what the diff --git a/src/group_theory/nilpotent.lean b/src/group_theory/nilpotent.lean index bea2a8d771b38..172bd1283a130 100644 --- a/src/group_theory/nilpotent.lean +++ b/src/group_theory/nilpotent.lean @@ -15,6 +15,9 @@ import tactic.tfae # Nilpotent groups +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + An API for nilpotent groups, that is, groups for which the upper central series reaches `⊤`. diff --git a/src/measure_theory/function/l1_space.lean b/src/measure_theory/function/l1_space.lean index c0716c5777ff1..30a1a57d68d80 100644 --- a/src/measure_theory/function/l1_space.lean +++ b/src/measure_theory/function/l1_space.lean @@ -9,6 +9,9 @@ import measure_theory.function.lp_order /-! # Integrable functions and `L¹` space +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In the first part of this file, the predicate `integrable` is defined and basic properties of integrable functions are proved. diff --git a/src/measure_theory/group/prod.lean b/src/measure_theory/group/prod.lean index 20da19bf2aed6..6e6bd774ca8db 100644 --- a/src/measure_theory/group/prod.lean +++ b/src/measure_theory/group/prod.lean @@ -8,6 +8,9 @@ import measure_theory.group.measure /-! # Measure theory in the product of groups + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. In this file we show properties about measure theory in products of measurable groups and properties of iterated integrals in measurable groups. diff --git a/src/measure_theory/integral/integrable_on.lean b/src/measure_theory/integral/integrable_on.lean index 41d7a83a9b1e8..63a9af38d5522 100644 --- a/src/measure_theory/integral/integrable_on.lean +++ b/src/measure_theory/integral/integrable_on.lean @@ -9,6 +9,9 @@ import analysis.normed_space.indicator_function /-! # Functions integrable on a set and at a filter +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define `integrable_on f s μ := integrable f (μ.restrict s)` and prove theorems like `integrable_on_union : integrable_on f (s ∪ t) μ ↔ integrable_on f s μ ∧ integrable_on f t μ`. diff --git a/src/measure_theory/measure/haar/basic.lean b/src/measure_theory/measure/haar/basic.lean index c988a723a3a6c..8765941a9aa26 100644 --- a/src/measure_theory/measure/haar/basic.lean +++ b/src/measure_theory/measure/haar/basic.lean @@ -11,6 +11,9 @@ import topology.algebra.group.compact /-! # Haar measure +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove the existence and uniqueness (up to scalar multiples) of Haar measure for a locally compact Hausdorff topological group. diff --git a/src/topology/homotopy/homotopy_group.lean b/src/topology/homotopy/homotopy_group.lean index 045f1d84ad6d6..8ca449bf538a4 100644 --- a/src/topology/homotopy/homotopy_group.lean +++ b/src/topology/homotopy/homotopy_group.lean @@ -9,6 +9,9 @@ import algebraic_topology.fundamental_groupoid.fundamental_group /-! # `n`th homotopy group +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define the `n`th homotopy group at `x`, `π n x`, as the equivalence classes of functions from the nth dimensional cube to the topological space `X` that send the boundary to the base point `x`, up to homotopic equivalence. diff --git a/src/topology/sheaves/local_predicate.lean b/src/topology/sheaves/local_predicate.lean index 4acca6972ccd2..04659ff17fd1b 100644 --- a/src/topology/sheaves/local_predicate.lean +++ b/src/topology/sheaves/local_predicate.lean @@ -11,6 +11,9 @@ import topology.sheaves.sheaf_condition.unique_gluing /-! # Functions satisfying a local predicate form a sheaf. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + At this stage, in `topology/sheaves/sheaf_of_functions.lean` we've proved that not-necessarily-continuous functions from a topological space into some type (or type family) form a sheaf. From 67237461d7cbf410706a6a06f4d40cbb594c32dc Mon Sep 17 00:00:00 2001 From: damiano Date: Thu, 1 Jun 2023 08:55:00 +0000 Subject: [PATCH 1126/1291] chore(field_theory/ratfunc): use `section` instead of `omit/include` (#19133) This hopefully closes [#4513](https://github.com/leanprover-community/mathlib4/issues/4513) (or at least mitigates it) and will settle [!4#4373](https://github.com/leanprover-community/mathlib4/pull/4373) in mathlib3. Co-authored-by: Anne Baanen --- src/field_theory/ratfunc.lean | 101 ++++++++++++++++++---------------- 1 file changed, 55 insertions(+), 46 deletions(-) diff --git a/src/field_theory/ratfunc.lean b/src/field_theory/ratfunc.lean index 07e44dd53cf0b..e56b902d45683 100644 --- a/src/field_theory/ratfunc.lean +++ b/src/field_theory/ratfunc.lean @@ -98,8 +98,7 @@ open_locale non_zero_divisors polynomial universes u v -variables (K : Type u) [hring : comm_ring K] [hdomain : is_domain K] -include hring +variable (K : Type u) /-- `ratfunc K` is `K(x)`, the field of rational functions over `K`. @@ -107,12 +106,15 @@ The inclusion of polynomials into `ratfunc` is `algebra_map K[X] (ratfunc K)`, the maps between `ratfunc K` and another field of fractions of `K[X]`, especially `fraction_ring K[X]`, are given by `is_localization.algebra_equiv`. -/ -structure ratfunc : Type u := of_fraction_ring :: +structure ratfunc [comm_ring K] : Type u := of_fraction_ring :: (to_fraction_ring : fraction_ring K[X]) namespace ratfunc -variables {K} +section comm_ring + +variable {K} +variable [comm_ring K] section rec @@ -159,7 +161,17 @@ begin exact localization.lift_on_mk _ _ _ _ end -include hdomain +lemma lift_on_condition_of_lift_on'_condition {P : Sort v} {f : ∀ (p q : K[X]), P} + (H : ∀ {p q a} (hq : q ≠ 0) (ha : a ≠ 0), f (a * p) (a * q) = f p q) + ⦃p q p' q' : K[X]⦄ (hq : q ≠ 0) (hq' : q' ≠ 0) (h : q' * p = q * p') : + f p q = f p' q' := +calc f p q = f (q' * p) (q' * q) : (H hq hq').symm + ... = f (q * p') (q * q') : by rw [h, mul_comm q'] + ... = f p' q' : H hq' hq + +section is_domain + +variable [is_domain K] /-- `ratfunc.mk (p q : K[X])` is `p / q` as a rational function. @@ -221,16 +233,6 @@ begin set_like.coe_mk] } end -omit hdomain -lemma lift_on_condition_of_lift_on'_condition {P : Sort v} {f : ∀ (p q : K[X]), P} - (H : ∀ {p q a} (hq : q ≠ 0) (ha : a ≠ 0), f (a * p) (a * q) = f p q) - ⦃p q p' q' : K[X]⦄ (hq : q ≠ 0) (hq' : q' ≠ 0) (h : q' * p = q * p') : - f p q = f p' q' := -calc f p q = f (q' * p) (q' * q) : (H hq hq').symm - ... = f (q * p') (q * q') : by rw [h, mul_comm q'] - ... = f p' q' : H hq' hq -include hdomain - /-- Non-dependent recursion principle for `ratfunc K`: if `f p q : P` for all `p q`, such that `f (a * p) (a * q) = f p q`, then we can find a value of `P` for all elements of `ratfunc K` by setting `lift_on' (p / q) f _ = f p q`. @@ -265,6 +267,8 @@ See also `induction_on`, which is a recursion principle defined in terms of `alg (λ ⟨p, q⟩, by simpa only [mk_coe_def, localization.mk_eq_mk'] using f p q (mem_non_zero_divisors_iff_ne_zero.mp q.2)) +end is_domain + end rec section field @@ -315,7 +319,9 @@ lemma of_fraction_ring_mul (p q : fraction_ring K[X]) : of_fraction_ring (p * q) = of_fraction_ring p * of_fraction_ring q := by unfold has_mul.mul ratfunc.mul -include hdomain +section is_domain + +variable [is_domain K] /-- Division of rational functions. -/ @[irreducible] protected def div : ratfunc K → ratfunc K → ratfunc K @@ -339,8 +345,9 @@ lemma mul_inv_cancel : ∀ {p : ratfunc K} (hp : p ≠ 0), p * p⁻¹ = 1 by simpa only [← of_fraction_ring_inv, ← of_fraction_ring_mul, ← of_fraction_ring_one] using _root_.mul_inv_cancel this +end is_domain + section has_smul -omit hdomain variables {R : Type*} @@ -372,10 +379,11 @@ begin { simp only } end -include hdomain +section is_domain + +variable [is_domain K] variables [monoid R] [distrib_mul_action R K[X]] -variables [htower : is_scalar_tower R K[X] K[X]] -include htower +variables [is_scalar_tower R K[X] K[X]] lemma mk_smul (c : R) (p q : K[X]) : ratfunc.mk (c • p) q = c • ratfunc.mk p q := @@ -389,12 +397,12 @@ end instance : is_scalar_tower R K[X] (ratfunc K) := ⟨λ c p q, q.induction_on' (λ q r _, by rw [← mk_smul, smul_assoc, mk_smul, mk_smul])⟩ +end is_domain + end has_smul variables (K) -omit hdomain - instance [subsingleton K] : subsingleton (ratfunc K) := to_fraction_ring_injective.subsingleton @@ -415,8 +423,9 @@ This is an auxiliary definition; `simp`-normal form is `is_localization.alg_equi map_add' := λ ⟨_⟩ ⟨_⟩, by simp [←of_fraction_ring_add], map_mul' := λ ⟨_⟩ ⟨_⟩, by simp [←of_fraction_ring_mul] } -omit hring +end field +section tactic_interlude -- pre-porting note: should comm_ring be disabled here? /-- Solve equations for `ratfunc K` by working in `fraction_ring K[X]`. -/ meta def frac_tac : tactic unit := `[repeat { rintro (⟨⟩ : ratfunc _) }, @@ -438,8 +447,8 @@ meta def smul_tac : tactic unit := localization.mk_zero, localization.add_mk_self, localization.neg_mk, of_fraction_ring_zero, ← of_fraction_ring_add, ← of_fraction_ring_neg]] -include hring - +end tactic_interlude +variable (K) instance : comm_ring (ratfunc K) := { add := (+), add_assoc := by frac_tac, @@ -473,7 +482,6 @@ variables {K} section lift_hom variables {G₀ L R S F : Type*} [comm_group_with_zero G₀] [field L] [comm_ring R] [comm_ring S] -omit hring /-- Lift a monoid homomorphism that maps polynomials `φ : R[X] →* S[X]` to a `ratfunc R →* ratfunc S`, @@ -636,10 +644,10 @@ lift_monoid_with_zero_hom_injective _ hφ end lift_hom -variables (K) -include hdomain +variable (K) -instance : field (ratfunc K) := +instance [is_domain K] : field (ratfunc K) := +by exact { inv := has_inv.inv, inv_zero := by frac_tac, div := (/), @@ -649,13 +657,12 @@ instance : field (ratfunc K) := .. ratfunc.comm_ring K, .. ratfunc.nontrivial K } -end field - section is_fraction_ring /-! ### `ratfunc` as field of fractions of `polynomial` -/ -include hdomain +section is_domain +variable [is_domain K] instance (R : Type*) [comm_semiring R] [algebra R K[X]] : algebra R (ratfunc K) := @@ -803,10 +810,6 @@ end lift_alg_hom variables (K) -omit hdomain - -include hdomain - /-- `ratfunc K` is the field of fractions of the polynomials over `K`. -/ instance : is_fraction_ring K[X] (ratfunc K) := { map_units := λ y, by rw ← of_fraction_ring_algebra_map; @@ -874,17 +877,21 @@ by simp only [localization.mk_eq_mk'_apply, of_fraction_ring_mk', is_localizatio by { ext x, simp [to_fraction_ring_ring_equiv, of_fraction_ring_eq, alg_equiv.coe_ring_equiv'] } +end is_domain + end is_fraction_ring +end comm_ring + +variable {K} + section num_denom /-! ### Numerator and denominator -/ open gcd_monoid polynomial -omit hring -variables [hfield : field K] -include hfield +variables [field K] /-- `ratfunc.num_denom` are numerator and denominator of a rational function over a field, normalized such that the denominator is monic. -/ @@ -1159,7 +1166,8 @@ section eval /-! ### Polynomial structure: `C`, `X`, `eval` -/ -include hdomain +section domain +variables [comm_ring K] [is_domain K] /-- `ratfunc.C a` is the constant rational function `a`. -/ def C : K →+* ratfunc K := @@ -1181,9 +1189,11 @@ def X : ratfunc K := algebra_map K[X] (ratfunc K) polynomial.X @[simp] lemma algebra_map_X : algebra_map K[X] (ratfunc K) polynomial.X = X := rfl -omit hring hdomain -variables [hfield : field K] -include hfield +end domain + +section field + +variables [field K] @[simp] lemma num_C (c : K) : num (C c) = polynomial.C c := num_algebra_map _ @@ -1274,14 +1284,14 @@ begin apply num_denom_mul, end +end field + end eval section int_degree open polynomial -omit hring - variables [field K] /-- `int_degree x` is the degree of the rational function `x`, defined as the difference between @@ -1365,7 +1375,6 @@ section laurent_series open power_series laurent_series hahn_series -omit hring variables {F : Type u} [field F] (p q : F[X]) (f g : ratfunc F) /-- The coercion `ratfunc F → laurent_series F` as bundled alg hom. -/ From 9d013ad8430ddddd350cff5c3db830278ded3c79 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 1 Jun 2023 13:20:10 +0000 Subject: [PATCH 1127/1291] doc(analysis/normed_space/pi_Lp): fix corrupt synchronization header (#19134) The missing blank line makes a mess in doc-gen. Hopefully this file was just strangely formatted, and this isn't a new bug in the script. --- src/analysis/normed_space/pi_Lp.lean | 1 + 1 file changed, 1 insertion(+) diff --git a/src/analysis/normed_space/pi_Lp.lean b/src/analysis/normed_space/pi_Lp.lean index f74d084aefc29..11056f463a979 100644 --- a/src/analysis/normed_space/pi_Lp.lean +++ b/src/analysis/normed_space/pi_Lp.lean @@ -12,6 +12,7 @@ import linear_algebra.matrix.basis > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. > Any changes to this file require a corresponding PR to mathlib4. + Given finitely many metric spaces, one can put the max distance on their product, but there is also a whole family of natural distances, indexed by a parameter `p : ℝ≥0∞`, that also induce the product topology. We define them in this file. For `0 < p < ∞`, the distance on `Π i, α i` From d35b4ff446f1421bd551fafa4b8efd98ac3ac408 Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Thu, 1 Jun 2023 16:18:47 +0000 Subject: [PATCH 1128/1291] chore(ring_theory): remove `splitting_field` dependency from `is_integrally_closed` (#19137) The only declarations involving splitting fields were `integral_closure.mem_lifts_of_monic_of_dvd_map` and `is_integrally_closed.eq_map_mul_C_of_dvd`. Both are similar to Gauss' lemma and their only use is to prove Gauss' lemma, so I moved them to that file. --- src/ring_theory/integrally_closed.lean | 64 ------------------- src/ring_theory/polynomial/gauss_lemma.lean | 68 +++++++++++++++++++++ 2 files changed, 68 insertions(+), 64 deletions(-) diff --git a/src/ring_theory/integrally_closed.lean b/src/ring_theory/integrally_closed.lean index a2483b48c5dc0..a2c2f43ab1a6c 100644 --- a/src/ring_theory/integrally_closed.lean +++ b/src/ring_theory/integrally_closed.lean @@ -3,7 +3,6 @@ Copyright (c) 2021 Anne Baanen. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Anne Baanen -/ -import field_theory.splitting_field import ring_theory.integral_closure import ring_theory.localization.integral @@ -21,8 +20,6 @@ integral over `R`. A special case of integrally closed domains are the Dedekind * `is_integrally_closed_iff K`, where `K` is a fraction field of `R`, states `R` is integrally closed iff it is the integral closure of `R` in `K` -* `eq_map_mul_C_of_dvd`: if `K = Frac(R)` and `g : K[X]` divides a monic polynomial with - coefficients in `R`, then `g * (C g.leading_coeff⁻¹)` has coefficients in `R` -/ open_locale non_zero_divisors polynomial @@ -132,29 +129,6 @@ open is_integrally_closed variables {R : Type*} [comm_ring R] variables (K : Type*) [field K] [algebra R K] -theorem mem_lifts_of_monic_of_dvd_map - {f : R[X]} (hf : f.monic) {g : K[X]} (hg : g.monic) (hd : g ∣ f.map (algebra_map R K)) : - g ∈ lifts (algebra_map (integral_closure R K) K) := -begin - haveI : is_scalar_tower R K g.splitting_field := splitting_field_aux.is_scalar_tower _ _ _, - have := mem_lift_of_splits_of_roots_mem_range (integral_closure R g.splitting_field) - ((splits_id_iff_splits _).2 $ splitting_field.splits g) (hg.map _) - (λ a ha, (set_like.ext_iff.mp (integral_closure R g.splitting_field).range_algebra_map _).mpr $ - roots_mem_integral_closure hf _), - { rw [lifts_iff_coeff_lifts, ←ring_hom.coe_range, subalgebra.range_algebra_map] at this, - refine (lifts_iff_coeff_lifts _).2 (λ n, _), - rw [← ring_hom.coe_range, subalgebra.range_algebra_map], - obtain ⟨p, hp, he⟩ := (set_like.mem_coe.mp (this n)), use [p, hp], - rw [is_scalar_tower.algebra_map_eq R K, coeff_map, ← eval₂_map, eval₂_at_apply] at he, - rw eval₂_eq_eval_map, apply (injective_iff_map_eq_zero _).1 _ _ he, - { apply ring_hom.injective } }, - rw [is_scalar_tower.algebra_map_eq R K _, ← map_map], - refine multiset.mem_of_le (roots.le_of_dvd ((hf.map _).map _).ne_zero _) ha, - { apply_instance }, - { exact map_dvd (algebra_map K g.splitting_field) hd }, - { apply splitting_field_aux.is_scalar_tower }, -end - variables [is_domain R] [is_fraction_ring R K] variables {L : Type*} [field L] [algebra K L] [algebra R L] [is_scalar_tower R K L] @@ -167,41 +141,3 @@ begin end end integral_closure - -namespace is_integrally_closed - -open integral_closure - -variables {R : Type*} [comm_ring R] [is_domain R] -variables (K : Type*) [field K] [algebra R K] [is_fraction_ring R K] - -/-- If `K = Frac(R)` and `g : K[X]` divides a monic polynomial with coefficients in `R`, then - `g * (C g.leading_coeff⁻¹)` has coefficients in `R` -/ -lemma eq_map_mul_C_of_dvd [is_integrally_closed R] {f : R[X]} (hf : f.monic) - {g : K[X]} (hg : g ∣ f.map (algebra_map R K)) : - ∃ g' : R[X], (g'.map (algebra_map R K)) * (C $ leading_coeff g) = g := -begin - have g_ne_0 : g ≠ 0 := ne_zero_of_dvd_ne_zero (monic.ne_zero $ hf.map (algebra_map R K)) hg, - suffices lem : ∃ g' : R[X], g'.map (algebra_map R K) = g * (C g.leading_coeff⁻¹), - { obtain ⟨g', hg'⟩ := lem, - use g', - rw [hg', mul_assoc, ← C_mul, inv_mul_cancel (leading_coeff_ne_zero.mpr g_ne_0), C_1, mul_one] }, - - have g_mul_dvd : g * (C g.leading_coeff⁻¹) ∣ f.map (algebra_map R K), - { rwa associated.dvd_iff_dvd_left (show associated (g * (C (g.leading_coeff⁻¹))) g, from _), - rw associated_mul_is_unit_left_iff, - exact is_unit_C.mpr (inv_ne_zero $ leading_coeff_ne_zero.mpr g_ne_0).is_unit }, - let algeq := (subalgebra.equiv_of_eq _ _ $ integral_closure_eq_bot R _).trans - (algebra.bot_equiv_of_injective $ is_fraction_ring.injective R $ K), - have : (algebra_map R _).comp algeq.to_alg_hom.to_ring_hom = - (integral_closure R _).to_subring.subtype, - { ext, conv_rhs { rw ← algeq.symm_apply_apply x }, refl }, - have H := ((mem_lifts _ ).1 - (mem_lifts_of_monic_of_dvd_map K hf (monic_mul_leading_coeff_inv g_ne_0) g_mul_dvd)), - refine ⟨map algeq.to_alg_hom.to_ring_hom _, _⟩, - use classical.some H, - rw [map_map, this], - exact classical.some_spec H -end - -end is_integrally_closed diff --git a/src/ring_theory/polynomial/gauss_lemma.lean b/src/ring_theory/polynomial/gauss_lemma.lean index 427d8cca5a60d..33888b6a063ed 100644 --- a/src/ring_theory/polynomial/gauss_lemma.lean +++ b/src/ring_theory/polynomial/gauss_lemma.lean @@ -3,6 +3,7 @@ Copyright (c) 2020 Aaron Anderson. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Aaron Anderson -/ +import field_theory.splitting_field import ring_theory.int.basic import ring_theory.localization.integral import ring_theory.integrally_closed @@ -14,6 +15,9 @@ import ring_theory.integrally_closed Gauss's Lemma is one of a few results pertaining to irreducibility of primitive polynomials. ## Main Results + - `is_integrally_closed.eq_map_mul_C_of_dvd`: if `R` is integrally closed, `K = Frac(R)` and + `g : K[X]` divides a monic polynomial with coefficients in `R`, then `g * (C g.leading_coeff⁻¹)` + has coefficients in `R` - `polynomial.monic.irreducible_iff_irreducible_map_fraction_map`: A monic polynomial over an integrally closed domain is irreducible iff it is irreducible in a fraction field @@ -35,6 +39,70 @@ open_locale non_zero_divisors polynomial variables {R : Type*} [comm_ring R] +section is_integrally_closed + +open polynomial +open integral_closure +open is_integrally_closed + +variables (K : Type*) [field K] [algebra R K] + +theorem integral_closure.mem_lifts_of_monic_of_dvd_map + {f : R[X]} (hf : f.monic) {g : K[X]} (hg : g.monic) (hd : g ∣ f.map (algebra_map R K)) : + g ∈ lifts (algebra_map (integral_closure R K) K) := +begin + haveI : is_scalar_tower R K g.splitting_field := splitting_field_aux.is_scalar_tower _ _ _, + have := mem_lift_of_splits_of_roots_mem_range (integral_closure R g.splitting_field) + ((splits_id_iff_splits _).2 $ splitting_field.splits g) (hg.map _) + (λ a ha, (set_like.ext_iff.mp (integral_closure R g.splitting_field).range_algebra_map _).mpr $ + roots_mem_integral_closure hf _), + { rw [lifts_iff_coeff_lifts, ←ring_hom.coe_range, subalgebra.range_algebra_map] at this, + refine (lifts_iff_coeff_lifts _).2 (λ n, _), + rw [← ring_hom.coe_range, subalgebra.range_algebra_map], + obtain ⟨p, hp, he⟩ := (set_like.mem_coe.mp (this n)), use [p, hp], + rw [is_scalar_tower.algebra_map_eq R K, coeff_map, ← eval₂_map, eval₂_at_apply] at he, + rw eval₂_eq_eval_map, apply (injective_iff_map_eq_zero _).1 _ _ he, + { apply ring_hom.injective } }, + rw [is_scalar_tower.algebra_map_eq R K _, ← map_map], + refine multiset.mem_of_le (roots.le_of_dvd ((hf.map _).map _).ne_zero _) ha, + { apply_instance }, + { exact map_dvd (algebra_map K g.splitting_field) hd }, + { apply splitting_field_aux.is_scalar_tower }, +end + +variables [is_domain R] [is_fraction_ring R K] + +/-- If `K = Frac(R)` and `g : K[X]` divides a monic polynomial with coefficients in `R`, then + `g * (C g.leading_coeff⁻¹)` has coefficients in `R` -/ +lemma is_integrally_closed.eq_map_mul_C_of_dvd [is_integrally_closed R] {f : R[X]} (hf : f.monic) + {g : K[X]} (hg : g ∣ f.map (algebra_map R K)) : + ∃ g' : R[X], (g'.map (algebra_map R K)) * (C $ leading_coeff g) = g := +begin + have g_ne_0 : g ≠ 0 := ne_zero_of_dvd_ne_zero (monic.ne_zero $ hf.map (algebra_map R K)) hg, + suffices lem : ∃ g' : R[X], g'.map (algebra_map R K) = g * (C g.leading_coeff⁻¹), + { obtain ⟨g', hg'⟩ := lem, + use g', + rw [hg', mul_assoc, ← C_mul, inv_mul_cancel (leading_coeff_ne_zero.mpr g_ne_0), C_1, mul_one] }, + + have g_mul_dvd : g * (C g.leading_coeff⁻¹) ∣ f.map (algebra_map R K), + { rwa associated.dvd_iff_dvd_left (show associated (g * (C (g.leading_coeff⁻¹))) g, from _), + rw associated_mul_is_unit_left_iff, + exact is_unit_C.mpr (inv_ne_zero $ leading_coeff_ne_zero.mpr g_ne_0).is_unit }, + let algeq := (subalgebra.equiv_of_eq _ _ $ integral_closure_eq_bot R _).trans + (algebra.bot_equiv_of_injective $ is_fraction_ring.injective R $ K), + have : (algebra_map R _).comp algeq.to_alg_hom.to_ring_hom = + (integral_closure R _).to_subring.subtype, + { ext, conv_rhs { rw ← algeq.symm_apply_apply x }, refl }, + have H := ((mem_lifts _).1 (integral_closure.mem_lifts_of_monic_of_dvd_map K hf + (monic_mul_leading_coeff_inv g_ne_0) g_mul_dvd)), + refine ⟨map algeq.to_alg_hom.to_ring_hom _, _⟩, + use classical.some H, + rw [map_map, this], + exact classical.some_spec H +end + +end is_integrally_closed + namespace polynomial section From b608348ffaeb7f557f2fd46876037abafd326ff3 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Thu, 1 Jun 2023 19:12:42 +0000 Subject: [PATCH 1129/1291] chore(ring_theory/derivation): split file (#19138) The Stone-Weierstrass theorem shouldn't require the definition of Lie algebras. Co-authored-by: Scott Morrison --- src/data/mv_polynomial/derivation.lean | 2 +- .../algebra/left_invariant_derivation.lean | 1 + src/geometry/manifold/derivation_bundle.lean | 2 +- .../basic.lean} | 142 +----------------- src/ring_theory/derivation/lie.lean | 50 ++++++ .../derivation/to_square_zero.lean | 121 +++++++++++++++ src/ring_theory/kaehler.lean | 2 +- 7 files changed, 179 insertions(+), 141 deletions(-) rename src/ring_theory/{derivation.lean => derivation/basic.lean} (68%) create mode 100644 src/ring_theory/derivation/lie.lean create mode 100644 src/ring_theory/derivation/to_square_zero.lean diff --git a/src/data/mv_polynomial/derivation.lean b/src/data/mv_polynomial/derivation.lean index 2e88c78ea8837..0591c3791fb13 100644 --- a/src/data/mv_polynomial/derivation.lean +++ b/src/data/mv_polynomial/derivation.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import data.mv_polynomial.supported -import ring_theory.derivation +import ring_theory.derivation.basic /-! # Derivations of multivariate polynomials diff --git a/src/geometry/manifold/algebra/left_invariant_derivation.lean b/src/geometry/manifold/algebra/left_invariant_derivation.lean index 0d3851627160c..6aab779498d14 100644 --- a/src/geometry/manifold/algebra/left_invariant_derivation.lean +++ b/src/geometry/manifold/algebra/left_invariant_derivation.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Nicolò Cavalleri -/ +import ring_theory.derivation.lie import geometry.manifold.derivation_bundle /-! diff --git a/src/geometry/manifold/derivation_bundle.lean b/src/geometry/manifold/derivation_bundle.lean index f463781aa276b..7a37e9acc559b 100644 --- a/src/geometry/manifold/derivation_bundle.lean +++ b/src/geometry/manifold/derivation_bundle.lean @@ -5,7 +5,7 @@ Authors: Nicolò Cavalleri -/ import geometry.manifold.algebra.smooth_functions -import ring_theory.derivation +import ring_theory.derivation.basic /-! diff --git a/src/ring_theory/derivation.lean b/src/ring_theory/derivation/basic.lean similarity index 68% rename from src/ring_theory/derivation.lean rename to src/ring_theory/derivation/basic.lean index 8816fcedc40de..09af046a4d1b2 100644 --- a/src/ring_theory/derivation.lean +++ b/src/ring_theory/derivation/basic.lean @@ -4,9 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Nicolò Cavalleri, Andrew Yang -/ -import algebra.lie.of_associative import ring_theory.adjoin.basic -import ring_theory.ideal.quotient_operations /-! # Derivations @@ -19,7 +17,11 @@ This file defines derivation. A derivation `D` from the `R`-algebra `A` to the ` - `derivation`: The type of `R`-derivations from `A` to `M`. This has an `A`-module structure. - `derivation.llcomp`: We may compose linear maps and derivations to obtain a derivation, and the composition is bilinear. + +See `ring_theory.derivation.lie` for - `derivation.lie_algebra`: The `R`-derivations from `A` to `A` form an lie algebra over `R`. + +and `ring_theory.derivation.to_square_zero` for - `derivation_to_square_zero_equiv_lift`: The `R`-derivations from `A` into a square-zero ideal `I` of `B` corresponds to the lifts `A →ₐ[R] B` of the map `A →ₐ[R] B ⧸ I`. @@ -341,142 +343,6 @@ coe_injective.add_comm_group _ coe_zero coe_add coe_neg coe_sub (λ _ _, rfl) ( end -section lie_structures - -/-! # Lie structures -/ - -variables (D : derivation R A A) {D1 D2 : derivation R A A} (r : R) (a b : A) - -/-- The commutator of derivations is again a derivation. -/ -instance : has_bracket (derivation R A A) (derivation R A A) := -⟨λ D1 D2, mk' (⁅(D1 : module.End R A), (D2 : module.End R A)⁆) $ λ a b, - by { simp only [ring.lie_def, map_add, id.smul_eq_mul, linear_map.mul_apply, leibniz, coe_fn_coe, - linear_map.sub_apply], ring, }⟩ - -@[simp] lemma commutator_coe_linear_map : - ↑⁅D1, D2⁆ = ⁅(D1 : module.End R A), (D2 : module.End R A)⁆ := rfl - -lemma commutator_apply : ⁅D1, D2⁆ a = D1 (D2 a) - D2 (D1 a) := rfl - -instance : lie_ring (derivation R A A) := -{ add_lie := λ d e f, by { ext a, simp only [commutator_apply, add_apply, map_add], ring, }, - lie_add := λ d e f, by { ext a, simp only [commutator_apply, add_apply, map_add], ring, }, - lie_self := λ d, by { ext a, simp only [commutator_apply, add_apply, map_add], ring_nf, }, - leibniz_lie := λ d e f, - by { ext a, simp only [commutator_apply, add_apply, sub_apply, map_sub], ring, } } - -instance : lie_algebra R (derivation R A A) := -{ lie_smul := λ r d e, by { ext a, simp only [commutator_apply, map_smul, smul_sub, smul_apply]}, - ..derivation.module } - -end lie_structures - end end derivation - -section to_square_zero - -universes u v w - -variables {R : Type u} {A : Type v} {B : Type w} [comm_semiring R] [comm_semiring A] [comm_ring B] -variables [algebra R A] [algebra R B] (I : ideal B) (hI : I ^ 2 = ⊥) - -/-- If `f₁ f₂ : A →ₐ[R] B` are two lifts of the same `A →ₐ[R] B ⧸ I`, - we may define a map `f₁ - f₂ : A →ₗ[R] I`. -/ -def diff_to_ideal_of_quotient_comp_eq (f₁ f₂ : A →ₐ[R] B) - (e : (ideal.quotient.mkₐ R I).comp f₁ = (ideal.quotient.mkₐ R I).comp f₂) : - A →ₗ[R] I := -linear_map.cod_restrict (I.restrict_scalars _) (f₁.to_linear_map - f₂.to_linear_map) -begin - intro x, - change f₁ x - f₂ x ∈ I, - rw [← ideal.quotient.eq, ← ideal.quotient.mkₐ_eq_mk R, ← alg_hom.comp_apply, e], - refl, -end - -@[simp] -lemma diff_to_ideal_of_quotient_comp_eq_apply (f₁ f₂ : A →ₐ[R] B) - (e : (ideal.quotient.mkₐ R I).comp f₁ = (ideal.quotient.mkₐ R I).comp f₂) (x : A) : - ((diff_to_ideal_of_quotient_comp_eq I f₁ f₂ e) x : B) = f₁ x - f₂ x := -rfl - -variables [algebra A B] [is_scalar_tower R A B] - -include hI - -/-- Given a tower of algebras `R → A → B`, and a square-zero `I : ideal B`, each lift `A →ₐ[R] B` -of the canonical map `A →ₐ[R] B ⧸ I` corresponds to a `R`-derivation from `A` to `I`. -/ -def derivation_to_square_zero_of_lift - (f : A →ₐ[R] B) (e : (ideal.quotient.mkₐ R I).comp f = is_scalar_tower.to_alg_hom R A (B ⧸ I)) : - derivation R A I := -begin - refine - { map_one_eq_zero' := _, - leibniz' := _, - ..(diff_to_ideal_of_quotient_comp_eq I f (is_scalar_tower.to_alg_hom R A B) _) }, - { rw e, ext, refl }, - { ext, change f 1 - algebra_map A B 1 = 0, rw [map_one, map_one, sub_self] }, - { intros x y, - let F := diff_to_ideal_of_quotient_comp_eq I f (is_scalar_tower.to_alg_hom R A B) - (by { rw e, ext, refl }), - have : (f x - algebra_map A B x) * (f y - algebra_map A B y) = 0, - { rw [← ideal.mem_bot, ← hI, pow_two], - convert (ideal.mul_mem_mul (F x).2 (F y).2) using 1 }, - ext, - dsimp only [submodule.coe_add, submodule.coe_mk, linear_map.coe_mk, - diff_to_ideal_of_quotient_comp_eq_apply, submodule.coe_smul_of_tower, - is_scalar_tower.coe_to_alg_hom', linear_map.to_fun_eq_coe], - simp only [map_mul, sub_mul, mul_sub, algebra.smul_def] at ⊢ this, - rw [sub_eq_iff_eq_add, sub_eq_iff_eq_add] at this, - rw this, - ring } -end - -lemma derivation_to_square_zero_of_lift_apply (f : A →ₐ[R] B) - (e : (ideal.quotient.mkₐ R I).comp f = is_scalar_tower.to_alg_hom R A (B ⧸ I)) - (x : A) : (derivation_to_square_zero_of_lift I hI f e x : B) = f x - algebra_map A B x := rfl - -/-- Given a tower of algebras `R → A → B`, and a square-zero `I : ideal B`, each `R`-derivation -from `A` to `I` corresponds to a lift `A →ₐ[R] B` of the canonical map `A →ₐ[R] B ⧸ I`. -/ -@[simps {attrs := []}] -def lift_of_derivation_to_square_zero (f : derivation R A I) : - A →ₐ[R] B := -{ to_fun := λ x, f x + algebra_map A B x, - map_one' := by rw [map_one, f.map_one_eq_zero, submodule.coe_zero, zero_add], - map_mul' := λ x y, begin - have : (f x : B) * (f y) = 0, - { rw [← ideal.mem_bot, ← hI, pow_two], convert (ideal.mul_mem_mul (f x).2 (f y).2) using 1 }, - simp only [map_mul, f.leibniz, add_mul, mul_add, submodule.coe_add, - submodule.coe_smul_of_tower, algebra.smul_def, this], - ring - end, - commutes' := λ r, - by simp only [derivation.map_algebra_map, eq_self_iff_true, zero_add, submodule.coe_zero, - ←is_scalar_tower.algebra_map_apply R A B r], - map_zero' := ((I.restrict_scalars R).subtype.comp f.to_linear_map + - (is_scalar_tower.to_alg_hom R A B).to_linear_map).map_zero, - ..((I.restrict_scalars R).subtype.comp f.to_linear_map + - (is_scalar_tower.to_alg_hom R A B).to_linear_map : A →ₗ[R] B) } - -@[simp] lemma lift_of_derivation_to_square_zero_mk_apply (d : derivation R A I) (x : A) : - ideal.quotient.mk I (lift_of_derivation_to_square_zero I hI d x) = algebra_map A (B ⧸ I) x := -by { rw [lift_of_derivation_to_square_zero_apply, map_add, - ideal.quotient.eq_zero_iff_mem.mpr (d x).prop, zero_add], refl } - -/-- Given a tower of algebras `R → A → B`, and a square-zero `I : ideal B`, -there is a 1-1 correspondance between `R`-derivations from `A` to `I` and -lifts `A →ₐ[R] B` of the canonical map `A →ₐ[R] B ⧸ I`. -/ -@[simps] -def derivation_to_square_zero_equiv_lift : - derivation R A I ≃ - { f : A →ₐ[R] B // (ideal.quotient.mkₐ R I).comp f = is_scalar_tower.to_alg_hom R A (B ⧸ I) } := -begin - refine ⟨λ d, ⟨lift_of_derivation_to_square_zero I hI d, _⟩, λ f, - (derivation_to_square_zero_of_lift I hI f.1 f.2 : _), _, _⟩, - { ext x, exact lift_of_derivation_to_square_zero_mk_apply I hI d x }, - { intro d, ext x, exact add_sub_cancel (d x : B) (algebra_map A B x) }, - { rintro ⟨f, hf⟩, ext x, exact sub_add_cancel (f x) (algebra_map A B x) } -end - -end to_square_zero diff --git a/src/ring_theory/derivation/lie.lean b/src/ring_theory/derivation/lie.lean new file mode 100644 index 0000000000000..fea858509675e --- /dev/null +++ b/src/ring_theory/derivation/lie.lean @@ -0,0 +1,50 @@ +/- +Copyright © 2020 Nicolò Cavalleri. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Nicolò Cavalleri, Andrew Yang +-/ +import algebra.lie.of_associative +import ring_theory.derivation.basic + +/-! +# Results + +- `derivation.lie_algebra`: The `R`-derivations from `A` to `A` form an lie algebra over `R`. + +-/ + +namespace derivation + +variables {R : Type*} [comm_ring R] +variables {A : Type*} [comm_ring A] [algebra R A] +variables (D : derivation R A A) {D1 D2 : derivation R A A} (a : A) + +section lie_structures + +/-! # Lie structures -/ + +/-- The commutator of derivations is again a derivation. -/ +instance : has_bracket (derivation R A A) (derivation R A A) := +⟨λ D1 D2, mk' (⁅(D1 : module.End R A), (D2 : module.End R A)⁆) $ λ a b, + by { simp only [ring.lie_def, map_add, algebra.id.smul_eq_mul, linear_map.mul_apply, leibniz, + coe_fn_coe, linear_map.sub_apply], ring, }⟩ + +@[simp] lemma commutator_coe_linear_map : + ↑⁅D1, D2⁆ = ⁅(D1 : module.End R A), (D2 : module.End R A)⁆ := rfl + +lemma commutator_apply : ⁅D1, D2⁆ a = D1 (D2 a) - D2 (D1 a) := rfl + +instance : lie_ring (derivation R A A) := +{ add_lie := λ d e f, by { ext a, simp only [commutator_apply, add_apply, map_add], ring, }, + lie_add := λ d e f, by { ext a, simp only [commutator_apply, add_apply, map_add], ring, }, + lie_self := λ d, by { ext a, simp only [commutator_apply, add_apply, map_add], ring_nf, }, + leibniz_lie := λ d e f, + by { ext a, simp only [commutator_apply, add_apply, sub_apply, map_sub], ring, } } + +instance : lie_algebra R (derivation R A A) := +{ lie_smul := λ r d e, by { ext a, simp only [commutator_apply, map_smul, smul_sub, smul_apply]}, + ..derivation.module } + +end lie_structures + +end derivation diff --git a/src/ring_theory/derivation/to_square_zero.lean b/src/ring_theory/derivation/to_square_zero.lean new file mode 100644 index 0000000000000..80791e2b871da --- /dev/null +++ b/src/ring_theory/derivation/to_square_zero.lean @@ -0,0 +1,121 @@ +/- +Copyright © 2020 Nicolò Cavalleri. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Nicolò Cavalleri, Andrew Yang +-/ +import ring_theory.derivation.basic +import ring_theory.ideal.quotient_operations + +/-! +# Results + +- `derivation_to_square_zero_equiv_lift`: The `R`-derivations from `A` into a square-zero ideal `I` + of `B` corresponds to the lifts `A →ₐ[R] B` of the map `A →ₐ[R] B ⧸ I`. + +-/ + +section to_square_zero + +universes u v w + +variables {R : Type u} {A : Type v} {B : Type w} [comm_semiring R] [comm_semiring A] [comm_ring B] +variables [algebra R A] [algebra R B] (I : ideal B) (hI : I ^ 2 = ⊥) + +/-- If `f₁ f₂ : A →ₐ[R] B` are two lifts of the same `A →ₐ[R] B ⧸ I`, + we may define a map `f₁ - f₂ : A →ₗ[R] I`. -/ +def diff_to_ideal_of_quotient_comp_eq (f₁ f₂ : A →ₐ[R] B) + (e : (ideal.quotient.mkₐ R I).comp f₁ = (ideal.quotient.mkₐ R I).comp f₂) : + A →ₗ[R] I := +linear_map.cod_restrict (I.restrict_scalars _) (f₁.to_linear_map - f₂.to_linear_map) +begin + intro x, + change f₁ x - f₂ x ∈ I, + rw [← ideal.quotient.eq, ← ideal.quotient.mkₐ_eq_mk R, ← alg_hom.comp_apply, e], + refl, +end + +@[simp] +lemma diff_to_ideal_of_quotient_comp_eq_apply (f₁ f₂ : A →ₐ[R] B) + (e : (ideal.quotient.mkₐ R I).comp f₁ = (ideal.quotient.mkₐ R I).comp f₂) (x : A) : + ((diff_to_ideal_of_quotient_comp_eq I f₁ f₂ e) x : B) = f₁ x - f₂ x := +rfl + +variables [algebra A B] [is_scalar_tower R A B] + +include hI + +/-- Given a tower of algebras `R → A → B`, and a square-zero `I : ideal B`, each lift `A →ₐ[R] B` +of the canonical map `A →ₐ[R] B ⧸ I` corresponds to a `R`-derivation from `A` to `I`. -/ +def derivation_to_square_zero_of_lift + (f : A →ₐ[R] B) (e : (ideal.quotient.mkₐ R I).comp f = is_scalar_tower.to_alg_hom R A (B ⧸ I)) : + derivation R A I := +begin + refine + { map_one_eq_zero' := _, + leibniz' := _, + ..(diff_to_ideal_of_quotient_comp_eq I f (is_scalar_tower.to_alg_hom R A B) _) }, + { rw e, ext, refl }, + { ext, change f 1 - algebra_map A B 1 = 0, rw [map_one, map_one, sub_self] }, + { intros x y, + let F := diff_to_ideal_of_quotient_comp_eq I f (is_scalar_tower.to_alg_hom R A B) + (by { rw e, ext, refl }), + have : (f x - algebra_map A B x) * (f y - algebra_map A B y) = 0, + { rw [← ideal.mem_bot, ← hI, pow_two], + convert (ideal.mul_mem_mul (F x).2 (F y).2) using 1 }, + ext, + dsimp only [submodule.coe_add, submodule.coe_mk, linear_map.coe_mk, + diff_to_ideal_of_quotient_comp_eq_apply, submodule.coe_smul_of_tower, + is_scalar_tower.coe_to_alg_hom', linear_map.to_fun_eq_coe], + simp only [map_mul, sub_mul, mul_sub, algebra.smul_def] at ⊢ this, + rw [sub_eq_iff_eq_add, sub_eq_iff_eq_add] at this, + rw this, + ring } +end + +lemma derivation_to_square_zero_of_lift_apply (f : A →ₐ[R] B) + (e : (ideal.quotient.mkₐ R I).comp f = is_scalar_tower.to_alg_hom R A (B ⧸ I)) + (x : A) : (derivation_to_square_zero_of_lift I hI f e x : B) = f x - algebra_map A B x := rfl + +/-- Given a tower of algebras `R → A → B`, and a square-zero `I : ideal B`, each `R`-derivation +from `A` to `I` corresponds to a lift `A →ₐ[R] B` of the canonical map `A →ₐ[R] B ⧸ I`. -/ +@[simps {attrs := []}] +def lift_of_derivation_to_square_zero (f : derivation R A I) : + A →ₐ[R] B := +{ to_fun := λ x, f x + algebra_map A B x, + map_one' := by rw [map_one, f.map_one_eq_zero, submodule.coe_zero, zero_add], + map_mul' := λ x y, begin + have : (f x : B) * (f y) = 0, + { rw [← ideal.mem_bot, ← hI, pow_two], convert (ideal.mul_mem_mul (f x).2 (f y).2) using 1 }, + simp only [map_mul, f.leibniz, add_mul, mul_add, submodule.coe_add, + submodule.coe_smul_of_tower, algebra.smul_def, this], + ring + end, + commutes' := λ r, + by simp only [derivation.map_algebra_map, eq_self_iff_true, zero_add, submodule.coe_zero, + ←is_scalar_tower.algebra_map_apply R A B r], + map_zero' := ((I.restrict_scalars R).subtype.comp f.to_linear_map + + (is_scalar_tower.to_alg_hom R A B).to_linear_map).map_zero, + ..((I.restrict_scalars R).subtype.comp f.to_linear_map + + (is_scalar_tower.to_alg_hom R A B).to_linear_map : A →ₗ[R] B) } + +@[simp] lemma lift_of_derivation_to_square_zero_mk_apply (d : derivation R A I) (x : A) : + ideal.quotient.mk I (lift_of_derivation_to_square_zero I hI d x) = algebra_map A (B ⧸ I) x := +by { rw [lift_of_derivation_to_square_zero_apply, map_add, + ideal.quotient.eq_zero_iff_mem.mpr (d x).prop, zero_add], refl } + +/-- Given a tower of algebras `R → A → B`, and a square-zero `I : ideal B`, +there is a 1-1 correspondance between `R`-derivations from `A` to `I` and +lifts `A →ₐ[R] B` of the canonical map `A →ₐ[R] B ⧸ I`. -/ +@[simps] +def derivation_to_square_zero_equiv_lift : + derivation R A I ≃ + { f : A →ₐ[R] B // (ideal.quotient.mkₐ R I).comp f = is_scalar_tower.to_alg_hom R A (B ⧸ I) } := +begin + refine ⟨λ d, ⟨lift_of_derivation_to_square_zero I hI d, _⟩, λ f, + (derivation_to_square_zero_of_lift I hI f.1 f.2 : _), _, _⟩, + { ext x, exact lift_of_derivation_to_square_zero_mk_apply I hI d x }, + { intro d, ext x, exact add_sub_cancel (d x : B) (algebra_map A B x) }, + { rintro ⟨f, hf⟩, ext x, exact sub_add_cancel (f x) (algebra_map A B x) } +end + +end to_square_zero diff --git a/src/ring_theory/kaehler.lean b/src/ring_theory/kaehler.lean index 8e910ef0c9657..14eedac487b93 100644 --- a/src/ring_theory/kaehler.lean +++ b/src/ring_theory/kaehler.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Nicolò Cavalleri, Andrew Yang -/ -import ring_theory.derivation +import ring_theory.derivation.to_square_zero import ring_theory.ideal.cotangent import ring_theory.is_tensor_product From 34ebaffc1d1e8e783fc05438ec2e70af87275ac9 Mon Sep 17 00:00:00 2001 From: Michael Stoll Date: Thu, 1 Jun 2023 20:20:54 +0000 Subject: [PATCH 1130/1291] feat(number_theory/sum_two_squares): add result for general n (#19054) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR resolves the "Todo" that was left in `number_theory.sum_two_squares`, i.e., it adds the result that characterizes natural numbers that are sums of two squares in terms of their prime factorization: ```lean lemma nat.eq_sq_add_sq_iff {n : ℕ} : (∃ x y : ℕ, n = x ^ 2 + y ^ 2) ↔ ∀ {q : ℕ} (hq : q.prime) (h : q % 4 = 3), even (padic_val_nat q n) := ... ``` There is some discussion on [Zulip](https://leanprover.zulipchat.com/#narrow/stream/217875-Is-there-code-for-X.3F/topic/Preferred.20spelling.20of.20p-adic.20valuation.3F/near/359896021) regarding how to spell out the condition on the right hand side. Co-authored-by: Michael Stoll <99838730+MichaelStollBayreuth@users.noreply.github.com> --- src/number_theory/sum_two_squares.lean | 234 +++++++++++++++++++++++-- 1 file changed, 224 insertions(+), 10 deletions(-) diff --git a/src/number_theory/sum_two_squares.lean b/src/number_theory/sum_two_squares.lean index 098a5058731c6..f949ca500609d 100644 --- a/src/number_theory/sum_two_squares.lean +++ b/src/number_theory/sum_two_squares.lean @@ -1,31 +1,245 @@ /- Copyright (c) 2019 Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Chris Hughes +Authors: Chris Hughes, Michael Stoll -/ import number_theory.zsqrtd.gaussian_int +import tactic.linear_combination /-! # Sums of two squares -Proof of Fermat's theorem on the sum of two squares. Every prime congruent to 1 mod 4 is the sum -of two squares. +Fermat's theorem on the sum of two squares. Every prime `p` congruent to 1 mod 4 is the +sum of two squares; see `nat.prime.sq_add_sq` (which has the weaker assumption `p % 4 ≠ 3`). -# Todo +We also give the result that characterizes the (positive) natural numbers that are sums +of two squares as those numbers `n` such that for every prime `q` congruent to 3 mod 4, the +exponent of the largest power of `q` dividing `n` is even; see `nat.eq_sq_add_sq_iff`. -Fully characterize the natural numbers that are the sum of two squares: those such that for every -prime p congruent to 3 mod 4, the largest power of p dividing them is even. +There is an alternative characterization as the numbers of the form `a^2 * b`, where `b` is a +natural number such that `-1` is a square modulo `b`; see `nat.eq_sq_add_sq_iff_eq_sq_mul`. -/ +section Fermat + open gaussian_int -/-- **Fermat's theorem on the sum of two squares**. Every prime congruent to 1 mod 4 is the sum +/-- **Fermat's theorem on the sum of two squares**. Every prime not congruent to 3 mod 4 is the sum of two squares. Also known as **Fermat's Christmas theorem**. -/ -lemma nat.prime.sq_add_sq {p : ℕ} [fact p.prime] (hp : p % 4 = 1) : +theorem nat.prime.sq_add_sq {p : ℕ} [fact p.prime] (hp : p % 4 ≠ 3) : ∃ a b : ℕ, a ^ 2 + b ^ 2 = p := begin apply sq_add_sq_of_nat_prime_of_not_irreducible p, - rw [principal_ideal_ring.irreducible_iff_prime, prime_iff_mod_four_eq_three_of_nat_prime p, hp], - norm_num + rwa [principal_ideal_ring.irreducible_iff_prime, prime_iff_mod_four_eq_three_of_nat_prime p], +end + +end Fermat + +/-! +### Generalities on sums of two squares +-/ + +section general + +/-- The set of sums of two squares is closed under multiplication in any commutative ring. +See also `sq_add_sq_mul_sq_add_sq`. -/ +lemma sq_add_sq_mul {R} [comm_ring R] {a b x y u v : R} (ha : a = x ^ 2 + y ^ 2) + (hb : b = u ^ 2 + v ^ 2) : ∃ r s : R, a * b = r ^ 2 + s ^ 2 := +⟨x * u - y * v, x * v + y * u, by {rw [ha, hb], ring}⟩ + +/-- The set of natural numbers that are sums of two squares is closed under multiplication. -/ +lemma nat.sq_add_sq_mul {a b x y u v : ℕ} (ha : a = x ^ 2 + y ^ 2) (hb : b = u ^ 2 + v ^ 2) : + ∃ r s : ℕ, a * b = r ^ 2 + s ^ 2 := +begin + zify at ha hb ⊢, + obtain ⟨r, s, h⟩ := sq_add_sq_mul ha hb, + refine ⟨r.nat_abs, s.nat_abs, _⟩, + simpa only [int.coe_nat_abs, sq_abs], +end + +end general + +/-! +### Results on when -1 is a square modulo a natural number +-/ + +section neg_one_square + +/-- If `-1` is a square modulo `n` and `m` divides `n`, then `-1` is also a square modulo `m`. -/ +-- This could be formulated for a general integer `a` in place of `-1`, +-- but it would not directly specialize to `-1`, +-- because `((-1 : ℤ) : zmod n)` is not the same as `(-1 : zmod n)`. +lemma zmod.is_square_neg_one_of_dvd {m n : ℕ} (hd : m ∣ n) (hs : is_square (-1 : zmod n)) : + is_square (-1 : zmod m) := +begin + let f : zmod n →+* zmod m := zmod.cast_hom hd _, + rw [← ring_hom.map_one f, ← ring_hom.map_neg], + exact hs.map f, +end + +/-- If `-1` is a square modulo coprime natural numbers `m` and `n`, then `-1` is also +a square modulo `m*n`. -/ +lemma zmod.is_square_neg_one_mul {m n : ℕ} (hc : m.coprime n) (hm : is_square (-1 : zmod m)) + (hn : is_square (-1 : zmod n)) : is_square (-1 : zmod (m * n)) := +begin + have : is_square (-1 : (zmod m) × (zmod n)), + { rw show (-1 : (zmod m) × (zmod n)) = ((-1 : zmod m), (-1 : zmod n)), from rfl, + obtain ⟨x, hx⟩ := hm, + obtain ⟨y, hy⟩ := hn, + rw [hx, hy], + exact ⟨(x, y), rfl⟩, }, + simpa only [ring_equiv.map_neg_one] using this.map (zmod.chinese_remainder hc).symm, +end + +/-- If a prime `p` divides `n` such that `-1` is a square modulo `n`, then `p % 4 ≠ 3`. -/ +lemma nat.prime.mod_four_ne_three_of_dvd_is_square_neg_one {p n : ℕ} (hpp : p.prime) (hp : p ∣ n) + (hs : is_square (-1 : zmod n)) : p % 4 ≠ 3 := +begin + obtain ⟨y, h⟩ := zmod.is_square_neg_one_of_dvd hp hs, + rw [← sq, eq_comm, show (-1 : zmod p) = -1 ^ 2, from by ring] at h, + haveI : fact p.prime := ⟨hpp⟩, + exact zmod.mod_four_ne_three_of_sq_eq_neg_sq' one_ne_zero h, +end + +/-- If `n` is a squarefree natural number, then `-1` is a square modulo `n` if and only if +`n` is not divisible by a prime `q` such that `q % 4 = 3`. -/ +lemma zmod.is_square_neg_one_iff {n : ℕ} (hn : squarefree n) : + is_square (-1 : zmod n) ↔ ∀ {q : ℕ}, q.prime → q ∣ n → q % 4 ≠ 3 := +begin + refine ⟨λ H q hqp hqd, hqp.mod_four_ne_three_of_dvd_is_square_neg_one hqd H, λ H, _⟩, + induction n using induction_on_primes with p n hpp ih, + { exact false.elim (hn.ne_zero rfl), }, + { exact ⟨0, by simp only [fin.zero_mul, neg_eq_zero, fin.one_eq_zero_iff]⟩, }, + { haveI : fact p.prime := ⟨hpp⟩, + have hcp : p.coprime n, + { by_contra hc, + exact hpp.not_unit (hn p $ mul_dvd_mul_left p $ hpp.dvd_iff_not_coprime.mpr hc), }, + have hp₁ := zmod.exists_sq_eq_neg_one_iff.mpr (H hpp (dvd_mul_right p n)), + exact zmod.is_square_neg_one_mul hcp hp₁ + (ih hn.of_mul_right (λ q hqp hqd, H hqp $ dvd_mul_of_dvd_right hqd _)), } end + +/-- If `n` is a squarefree natural number, then `-1` is a square modulo `n` if and only if +`n` has no divisor `q` that is `≡ 3 mod 4`. -/ +lemma zmod.is_square_neg_one_iff' {n : ℕ} (hn : squarefree n) : + is_square (-1 : zmod n) ↔ ∀ {q : ℕ}, q ∣ n → q % 4 ≠ 3 := +begin + have help : ∀ a b : zmod 4, a ≠ 3 → b ≠ 3 → a * b ≠ 3 := by dec_trivial, + rw zmod.is_square_neg_one_iff hn, + refine ⟨λ H, induction_on_primes _ _ (λ p q hp hq hpq, _), λ H q hq₁, H⟩, + { exact λ _, by norm_num, }, + { exact λ _, by norm_num, }, + { replace hp := H hp (dvd_of_mul_right_dvd hpq), + replace hq := hq (dvd_of_mul_left_dvd hpq), + rw [(show 3 = 3 % 4, from by norm_num), ne.def, ← zmod.nat_coe_eq_nat_coe_iff'] at hp hq ⊢, + rw nat.cast_mul, + exact help p q hp hq, } +end + +/-! +### Relation to sums of two squares +-/ + +/-- If `-1` is a square modulo the natural number `n`, then `n` is a sum of two squares. -/ +lemma nat.eq_sq_add_sq_of_is_square_mod_neg_one {n : ℕ} (h : is_square (-1 : zmod n)) : + ∃ x y : ℕ, n = x ^ 2 + y ^ 2 := +begin + induction n using induction_on_primes with p n hpp ih, + { exact ⟨0, 0, rfl⟩, }, + { exact ⟨0, 1, rfl⟩, }, + { haveI : fact p.prime := ⟨hpp⟩, + have hp : is_square (-1 : zmod p) := zmod.is_square_neg_one_of_dvd ⟨n, rfl⟩ h, + obtain ⟨u, v, huv⟩ := nat.prime.sq_add_sq (zmod.exists_sq_eq_neg_one_iff.mp hp), + obtain ⟨x, y, hxy⟩ := ih (zmod.is_square_neg_one_of_dvd ⟨p, mul_comm _ _⟩ h), + exact nat.sq_add_sq_mul huv.symm hxy, } +end + +/-- If the integer `n` is a sum of two squares of coprime integers, +then `-1` is a square modulo `n`. -/ +lemma zmod.is_square_neg_one_of_eq_sq_add_sq_of_is_coprime {n x y : ℤ} (h : n = x ^ 2 + y ^ 2) + (hc : is_coprime x y) : is_square (-1 : zmod n.nat_abs) := +begin + obtain ⟨u, v, huv⟩ : is_coprime x n, + { have hc2 : is_coprime (x ^ 2) (y ^ 2) := hc.pow, + rw show y ^ 2 = n + (-1) * x ^ 2, from by {rw h, ring} at hc2, + exact (is_coprime.pow_left_iff zero_lt_two).mp hc2.of_add_mul_right_right, }, + have H : (u * y) * (u * y) - (-1) = n * (-v ^ 2 * n + u ^ 2 + 2 * v) := + by linear_combination -u ^ 2 * h + (n * v - u * x - 1) * huv, + refine ⟨u * y, _⟩, + norm_cast, + rw (by push_cast : (-1 : zmod n.nat_abs) = (-1 : ℤ)), + exact (zmod.int_coe_eq_int_coe_iff_dvd_sub _ _ _).mpr (int.nat_abs_dvd.mpr ⟨_, H⟩), +end + +/-- If the natural number `n` is a sum of two squares of coprime natural numbers, then +`-1` is a square modulo `n`. -/ +lemma zmod.is_square_neg_one_of_eq_sq_add_sq_of_coprime {n x y : ℕ} (h : n = x ^ 2 + y ^ 2) + (hc : x.coprime y) : is_square (-1 : zmod n) := +begin + zify at *, + exact zmod.is_square_neg_one_of_eq_sq_add_sq_of_is_coprime h hc.is_coprime, +end + +/-- A natural number `n` is a sum of two squares if and only if `n = a^2 * b` with natural +numbers `a` and `b` such that `-1` is a square modulo `b`. -/ +lemma nat.eq_sq_add_sq_iff_eq_sq_mul {n : ℕ} : + (∃ x y : ℕ, n = x ^ 2 + y ^ 2) ↔ ∃ a b : ℕ, n = a ^ 2 * b ∧ is_square (-1 : zmod b) := +begin + split, + { rintros ⟨x, y, h⟩, + by_cases hxy : x = 0 ∧ y = 0, + { exact ⟨0, 1, by rw [h, hxy.1, hxy.2, zero_pow zero_lt_two, add_zero, zero_mul], + ⟨0, by rw [zero_mul, neg_eq_zero, fin.one_eq_zero_iff]⟩⟩, }, + { have hg := nat.pos_of_ne_zero (mt nat.gcd_eq_zero_iff.mp hxy), + obtain ⟨g, x₁, y₁, h₁, h₂, h₃, h₄⟩ := nat.exists_coprime' hg, + exact ⟨g, x₁ ^ 2 + y₁ ^ 2, by {rw [h, h₃, h₄], ring}, + zmod.is_square_neg_one_of_eq_sq_add_sq_of_coprime rfl h₂⟩, } }, + { rintros ⟨a, b, h₁, h₂⟩, + obtain ⟨x', y', h⟩ := nat.eq_sq_add_sq_of_is_square_mod_neg_one h₂, + exact ⟨a * x', a * y', by {rw [h₁, h], ring}⟩, } +end + +end neg_one_square + +/-! +### Characterization in terms of the prime factorization +-/ + +section main + +/-- A (positive) natural number `n` is a sum of two squares if and only if the exponent of +every prime `q` such that `q % 4 = 3` in the prime factorization of `n` is even. +(The assumption `0 < n` is not present, since for `n = 0`, both sides are satisfied; +the right hand side holds, since `padic_val_nat q 0 = 0` by definition.) -/ +lemma nat.eq_sq_add_sq_iff {n : ℕ} : + (∃ x y : ℕ, n = x ^ 2 + y ^ 2) ↔ ∀ {q : ℕ}, q.prime → q % 4 = 3 → even (padic_val_nat q n) := +begin + rcases n.eq_zero_or_pos with rfl | hn₀, + { exact ⟨λ H q hq h, (@padic_val_nat.zero q).symm ▸ even_zero, λ H, ⟨0, 0, rfl⟩⟩, }, + -- now `0 < n` + rw nat.eq_sq_add_sq_iff_eq_sq_mul, + refine ⟨λ H q hq h, _, λ H, _⟩, + { obtain ⟨a, b, h₁, h₂⟩ := H, + have hqb := padic_val_nat.eq_zero_of_not_dvd + (λ hf, (hq.mod_four_ne_three_of_dvd_is_square_neg_one hf h₂) h), + have hab : a ^ 2 * b ≠ 0 := h₁ ▸ hn₀.ne', + have ha₂ := left_ne_zero_of_mul hab, + have ha := mt sq_eq_zero_iff.mpr ha₂, + have hb := right_ne_zero_of_mul hab, + haveI hqi : fact q.prime := ⟨hq⟩, + simp_rw [h₁, padic_val_nat.mul ha₂ hb, padic_val_nat.pow 2 ha, hqb, add_zero], + exact even_two_mul _, }, + { obtain ⟨b, a, hb₀, ha₀, hab, hb⟩ := nat.sq_mul_squarefree_of_pos hn₀, + refine ⟨a, b, hab.symm, (zmod.is_square_neg_one_iff hb).mpr (λ q hqp hqb hq4, _)⟩, + refine nat.odd_iff_not_even.mp _ (H hqp hq4), + have hqb' : padic_val_nat q b = 1 := + b.factorization_def hqp ▸ le_antisymm (nat.squarefree.factorization_le_one _ hb) + ((hqp.dvd_iff_one_le_factorization hb₀.ne').mp hqb), + haveI hqi : fact q.prime := ⟨hqp⟩, + simp_rw [← hab, padic_val_nat.mul (pow_ne_zero 2 ha₀.ne') hb₀.ne', hqb', + padic_val_nat.pow 2 ha₀.ne'], + exact odd_two_mul_add_one _, } +end + +end main From 6a5c85000ab93fe5dcfdf620676f614ba8e18c26 Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Fri, 2 Jun 2023 04:32:37 +0000 Subject: [PATCH 1131/1291] chore(analysis/special_functions/exp_deriv): downgrade import (#19139) Move lemma `complex.is_open_map_exp` from `special_functions/exp_deriv` to right before its (unique) place of use, in `complex.exp_local_homeomorph` in `special_functions/log_deriv`. Morally these belong together: being an open map and being a local homeomorphism are closely tied, they are both consequences of the inverse function theorem and the point of both is to set up being able to differentiate complex log as the inverse function of complex exp. This removes the inverse function theorem from the dependencies of `special_functions/exp_deriv` (a surprisingly widely-imported file). --- src/analysis/special_functions/complex/log_deriv.lean | 4 ++++ src/analysis/special_functions/exp_deriv.lean | 4 ---- 2 files changed, 4 insertions(+), 4 deletions(-) diff --git a/src/analysis/special_functions/complex/log_deriv.lean b/src/analysis/special_functions/complex/log_deriv.lean index 2d43df3418d50..139d31bbdbe4a 100644 --- a/src/analysis/special_functions/complex/log_deriv.lean +++ b/src/analysis/special_functions/complex/log_deriv.lean @@ -3,6 +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, Benjamin Davidson -/ +import analysis.calculus.inverse import analysis.special_functions.complex.log import analysis.special_functions.exp_deriv @@ -19,6 +20,9 @@ open set filter open_locale real topology +lemma is_open_map_exp : is_open_map exp := +open_map_of_strict_deriv has_strict_deriv_at_exp exp_ne_zero + /-- `complex.exp` as a `local_homeomorph` with `source = {z | -π < im z < π}` and `target = {z | 0 < re z} ∪ {z | im z ≠ 0}`. This definition is used to prove that `complex.log` is complex differentiable at all points but the negative real semi-axis. -/ diff --git a/src/analysis/special_functions/exp_deriv.lean b/src/analysis/special_functions/exp_deriv.lean index ff0b7ade2a641..666baa6e15813 100644 --- a/src/analysis/special_functions/exp_deriv.lean +++ b/src/analysis/special_functions/exp_deriv.lean @@ -3,7 +3,6 @@ Copyright (c) 2018 Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes, Abhimanyu Pallavi Sudhir, Jean Lo, Calle Sönne -/ -import analysis.calculus.inverse import analysis.complex.real_deriv /-! @@ -67,9 +66,6 @@ lemma has_strict_fderiv_at_exp_real (x : ℂ) : has_strict_fderiv_at exp (exp x • (1 : ℂ →L[ℝ] ℂ)) x := (has_strict_deriv_at_exp x).complex_to_real_fderiv -lemma is_open_map_exp : is_open_map exp := -open_map_of_strict_deriv has_strict_deriv_at_exp exp_ne_zero - end complex section From a16665637b378379689c566204817ae792ac8b39 Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Fri, 2 Jun 2023 05:41:32 +0000 Subject: [PATCH 1132/1291] chore(analysis/convex/specific_functions/deriv): remove unnecessary imports (#19140) I accidentally left some extra imports in this file during the split #19031. --- src/analysis/convex/specific_functions/deriv.lean | 2 -- src/number_theory/bertrand.lean | 1 + 2 files changed, 1 insertion(+), 2 deletions(-) diff --git a/src/analysis/convex/specific_functions/deriv.lean b/src/analysis/convex/specific_functions/deriv.lean index 907264c088293..f8ed93d429deb 100644 --- a/src/analysis/convex/specific_functions/deriv.lean +++ b/src/analysis/convex/specific_functions/deriv.lean @@ -3,11 +3,9 @@ Copyright (c) 2020 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov, Sébastien Gouëzel -/ -import analysis.convex.specific_functions.basic import analysis.calculus.deriv.zpow import analysis.special_functions.pow.deriv import analysis.special_functions.sqrt -import tactic.linear_combination /-! # Collection of convex functions diff --git a/src/number_theory/bertrand.lean b/src/number_theory/bertrand.lean index a3cde2c5d0267..38e6d956438a7 100644 --- a/src/number_theory/bertrand.lean +++ b/src/number_theory/bertrand.lean @@ -6,6 +6,7 @@ Authors: Patrick Stevens, Bolton Bailey import data.nat.choose.factorization import data.nat.prime_norm_num import number_theory.primorial +import analysis.convex.specific_functions.basic import analysis.convex.specific_functions.deriv /-! From 2ebc1d6c2fed9f54c95bbc3998eaa5570527129a Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Fri, 2 Jun 2023 05:41:33 +0000 Subject: [PATCH 1133/1291] chore(*): add mathlib4 synchronization comments (#19141) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebraic_geometry.stalks` * `algebraic_geometry.structure_sheaf` * `analysis.analytic.basic` * `analysis.calculus.cont_diff_def` * `analysis.calculus.iterated_deriv` * `analysis.calculus.mean_value` * `analysis.inner_product_space.adjoint` * `analysis.inner_product_space.positive` * `analysis.normed_space.lp_equiv` * `analysis.normed_space.lp_space` * `measure_theory.constructions.pi` * `measure_theory.function.convergence_in_measure` * `measure_theory.function.locally_integrable` * `measure_theory.measure.haar.of_basis` * `probability.probability_mass_function.uniform` * `ring_theory.filtration` * `topology.algebra.valued_field` * `topology.metric_space.kuratowski` --- src/algebraic_geometry/stalks.lean | 3 +++ src/algebraic_geometry/structure_sheaf.lean | 3 +++ src/analysis/analytic/basic.lean | 3 +++ src/analysis/calculus/cont_diff_def.lean | 3 +++ src/analysis/calculus/iterated_deriv.lean | 3 +++ src/analysis/calculus/mean_value.lean | 3 +++ src/analysis/inner_product_space/adjoint.lean | 3 +++ src/analysis/inner_product_space/positive.lean | 3 +++ src/analysis/normed_space/lp_equiv.lean | 3 +++ src/analysis/normed_space/lp_space.lean | 3 +++ src/measure_theory/constructions/pi.lean | 3 +++ src/measure_theory/function/convergence_in_measure.lean | 3 +++ src/measure_theory/function/locally_integrable.lean | 3 +++ src/measure_theory/measure/haar/of_basis.lean | 3 +++ src/probability/probability_mass_function/uniform.lean | 3 +++ src/ring_theory/filtration.lean | 3 +++ src/topology/algebra/valued_field.lean | 3 +++ src/topology/metric_space/kuratowski.lean | 3 +++ 18 files changed, 54 insertions(+) diff --git a/src/algebraic_geometry/stalks.lean b/src/algebraic_geometry/stalks.lean index a702694af2196..4f5b48591d9ea 100644 --- a/src/algebraic_geometry/stalks.lean +++ b/src/algebraic_geometry/stalks.lean @@ -10,6 +10,9 @@ import topology.sheaves.stalks /-! # Stalks for presheaved spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file lifts constructions of stalks and pushforwards of stalks to work with the category of presheafed spaces. Additionally, we prove that restriction of presheafed spaces does not change the stalks. diff --git a/src/algebraic_geometry/structure_sheaf.lean b/src/algebraic_geometry/structure_sheaf.lean index a557011cfabb5..1f2848001ed88 100644 --- a/src/algebraic_geometry/structure_sheaf.lean +++ b/src/algebraic_geometry/structure_sheaf.lean @@ -13,6 +13,9 @@ import ring_theory.subring.basic /-! # The structure sheaf on `prime_spectrum R`. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define the structure sheaf on `Top.of (prime_spectrum R)`, for a commutative ring `R` and prove basic properties about it. We define this as a subsheaf of the sheaf of dependent functions into the localizations, cut out by the condition that the function must be locally equal to a ratio of diff --git a/src/analysis/analytic/basic.lean b/src/analysis/analytic/basic.lean index c33669722d9d5..ce82b42524bcd 100644 --- a/src/analysis/analytic/basic.lean +++ b/src/analysis/analytic/basic.lean @@ -11,6 +11,9 @@ import topology.algebra.infinite_sum.module /-! # Analytic functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A function is analytic in one dimension around `0` if it can be written as a converging power series `Σ pₙ zⁿ`. This definition can be extended to any dimension (even in infinite dimension) by requiring that `pₙ` is a continuous `n`-multilinear map. In general, `pₙ` is not unique (in two diff --git a/src/analysis/calculus/cont_diff_def.lean b/src/analysis/calculus/cont_diff_def.lean index 80e603f1b05f7..bd25338ef2ea3 100644 --- a/src/analysis/calculus/cont_diff_def.lean +++ b/src/analysis/calculus/cont_diff_def.lean @@ -12,6 +12,9 @@ import analysis.calculus.formal_multilinear_series /-! # Higher differentiability +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A function is `C^1` on a domain if it is differentiable there, and its derivative is continuous. By induction, it is `C^n` if it is `C^{n-1}` and its (n-1)-th derivative is `C^1` there or, equivalently, if it is `C^1` and its derivative is `C^{n-1}`. diff --git a/src/analysis/calculus/iterated_deriv.lean b/src/analysis/calculus/iterated_deriv.lean index 634a49a1ec745..1c1f1e4c36a40 100644 --- a/src/analysis/calculus/iterated_deriv.lean +++ b/src/analysis/calculus/iterated_deriv.lean @@ -9,6 +9,9 @@ import analysis.calculus.cont_diff_def /-! # One-dimensional iterated derivatives +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define the `n`-th derivative of a function `f : 𝕜 → F` as a function `iterated_deriv n f : 𝕜 → F`, as well as a version on domains `iterated_deriv_within n f s : 𝕜 → F`, and prove their basic properties. diff --git a/src/analysis/calculus/mean_value.lean b/src/analysis/calculus/mean_value.lean index 5b38db38f1e77..1fcb04ffbf97d 100644 --- a/src/analysis/calculus/mean_value.lean +++ b/src/analysis/calculus/mean_value.lean @@ -13,6 +13,9 @@ import topology.instances.real_vector_space /-! # The mean value inequality and equalities +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove the following facts: * `convex.norm_image_sub_le_of_norm_deriv_le` : if `f` is differentiable on a convex set `s` diff --git a/src/analysis/inner_product_space/adjoint.lean b/src/analysis/inner_product_space/adjoint.lean index 294543aa89761..768e5627ed4a4 100644 --- a/src/analysis/inner_product_space/adjoint.lean +++ b/src/analysis/inner_product_space/adjoint.lean @@ -10,6 +10,9 @@ import analysis.inner_product_space.pi_L2 /-! # Adjoint of operators on Hilbert spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given an operator `A : E →L[𝕜] F`, where `E` and `F` are Hilbert spaces, its adjoint `adjoint A : F →L[𝕜] E` is the unique operator such that `⟪x, A y⟫ = ⟪adjoint A x, y⟫` for all `x` and `y`. diff --git a/src/analysis/inner_product_space/positive.lean b/src/analysis/inner_product_space/positive.lean index 8b35a1d815b43..d9437594cc446 100644 --- a/src/analysis/inner_product_space/positive.lean +++ b/src/analysis/inner_product_space/positive.lean @@ -8,6 +8,9 @@ import analysis.inner_product_space.adjoint /-! # Positive operators +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define positive operators in a Hilbert space. We follow Bourbaki's choice of requiring self adjointness in the definition. diff --git a/src/analysis/normed_space/lp_equiv.lean b/src/analysis/normed_space/lp_equiv.lean index b705a4bd3ce8c..9555744026e36 100644 --- a/src/analysis/normed_space/lp_equiv.lean +++ b/src/analysis/normed_space/lp_equiv.lean @@ -10,6 +10,9 @@ import topology.continuous_function.bounded /-! # Equivalences among $L^p$ spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we collect a variety of equivalences among various $L^p$ spaces. In particular, when `α` is a `fintype`, given `E : α → Type u` and `p : ℝ≥0∞`, there is a natural linear isometric equivalence `lp_pi_Lpₗᵢ : lp E p ≃ₗᵢ pi_Lp p E`. In addition, when `α` is a discrete topological diff --git a/src/analysis/normed_space/lp_space.lean b/src/analysis/normed_space/lp_space.lean index 4d69acab9ca0d..519890f3f49d6 100644 --- a/src/analysis/normed_space/lp_space.lean +++ b/src/analysis/normed_space/lp_space.lean @@ -11,6 +11,9 @@ import topology.algebra.order.liminf_limsup /-! # ℓp space +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file describes properties of elements `f` of a pi-type `Π i, E i` with finite "norm", defined for `p:ℝ≥0∞` as the size of the support of `f` if `p=0`, `(∑' a, ‖f a‖^p) ^ (1/p)` for `0 < p < ∞` and `⨆ a, ‖f a‖` for `p=∞`. diff --git a/src/measure_theory/constructions/pi.lean b/src/measure_theory/constructions/pi.lean index 2b2c400417f35..a616ee6c52a46 100644 --- a/src/measure_theory/constructions/pi.lean +++ b/src/measure_theory/constructions/pi.lean @@ -10,6 +10,9 @@ import topology.constructions /-! # Product measures +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define and prove properties about finite products of measures (and at some point, countable products of measures). diff --git a/src/measure_theory/function/convergence_in_measure.lean b/src/measure_theory/function/convergence_in_measure.lean index ee8737e30c14c..823d8e060d4c4 100644 --- a/src/measure_theory/function/convergence_in_measure.lean +++ b/src/measure_theory/function/convergence_in_measure.lean @@ -10,6 +10,9 @@ import measure_theory.function.lp_space /-! # Convergence in measure +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define convergence in measure which is one of the many notions of convergence in probability. A sequence of functions `f` is said to converge in measure to some function `g` if for all `ε > 0`, the measure of the set `{x | ε ≤ dist (f i x) (g x)}` tends to 0 as `i` diff --git a/src/measure_theory/function/locally_integrable.lean b/src/measure_theory/function/locally_integrable.lean index af86545e34b22..a60ff6ae5fe66 100644 --- a/src/measure_theory/function/locally_integrable.lean +++ b/src/measure_theory/function/locally_integrable.lean @@ -8,6 +8,9 @@ import measure_theory.integral.integrable_on /-! # Locally integrable functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A function is called *locally integrable* (`measure_theory.locally_integrable`) if it is integrable on a neighborhood of every point. More generally, it is *locally integrable on `s`* if it is locally integrable on a neighbourhood within `s` of any point of `s`. diff --git a/src/measure_theory/measure/haar/of_basis.lean b/src/measure_theory/measure/haar/of_basis.lean index 0c0cf1c9bb18f..d8aed76d5d53d 100644 --- a/src/measure_theory/measure/haar/of_basis.lean +++ b/src/measure_theory/measure/haar/of_basis.lean @@ -9,6 +9,9 @@ import analysis.inner_product_space.pi_L2 /-! # Additive Haar measure constructed from a basis +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given a basis of a finite-dimensional real vector space, we define the corresponding Lebesgue measure, which gives measure `1` to the parallelepiped spanned by the basis. diff --git a/src/probability/probability_mass_function/uniform.lean b/src/probability/probability_mass_function/uniform.lean index 2bc42f9c97ed4..42fbd7618dc16 100644 --- a/src/probability/probability_mass_function/uniform.lean +++ b/src/probability/probability_mass_function/uniform.lean @@ -8,6 +8,9 @@ import probability.probability_mass_function.constructions /-! # Uniform Probability Mass Functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines a number of uniform `pmf` distributions from various inputs, uniformly drawing from the corresponding object. diff --git a/src/ring_theory/filtration.lean b/src/ring_theory/filtration.lean index 334bdab3e4de4..b037c0da93ecb 100644 --- a/src/ring_theory/filtration.lean +++ b/src/ring_theory/filtration.lean @@ -14,6 +14,9 @@ import order.hom.lattice # `I`-filtrations of modules +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains the definitions and basic results around (stable) `I`-filtrations of modules. ## Main results diff --git a/src/topology/algebra/valued_field.lean b/src/topology/algebra/valued_field.lean index 9bf06e431ca65..5d61bc738ed66 100644 --- a/src/topology/algebra/valued_field.lean +++ b/src/topology/algebra/valued_field.lean @@ -11,6 +11,9 @@ import topology.algebra.uniform_field /-! # Valued fields and their completions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we study the topology of a field `K` endowed with a valuation (in our application to adic spaces, `K` will be the valuation field associated to some valuation on a ring, defined in valuation.basic). diff --git a/src/topology/metric_space/kuratowski.lean b/src/topology/metric_space/kuratowski.lean index 99fd9607cca37..f5e0eba971407 100644 --- a/src/topology/metric_space/kuratowski.lean +++ b/src/topology/metric_space/kuratowski.lean @@ -9,6 +9,9 @@ import topology.sets.compacts /-! # The Kuratowski embedding +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Any separable metric space can be embedded isometrically in `ℓ^∞(ℝ)`. -/ From 69b2e97a276619372b19cf80fc1e91b05ae2baa4 Mon Sep 17 00:00:00 2001 From: Anatole Dedecker Date: Fri, 2 Jun 2023 09:01:24 +0000 Subject: [PATCH 1134/1291] chore(ring_theory/tensor_product): replace `is_scalar_tower` by `smul_comm_class` in `left_algebra` (#19118) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit With this change, this instance works for both restriction (e.g tensor product over `ℂ` of complex algebras is a real algebra) and extension (e.g tensor product over `ℝ` of complex algebras is a complex algebra) of scalars. [Zulip](https://leanprover.zulipchat.com/#narrow/stream/217875-Is-there-code-for-X.3F/topic/Tensor.20products) I also removes [algebra.tensor_product.tensor_product.is_scalar_tower](https://leanprover-community.github.io/mathlib_docs/ring_theory/tensor_product.html#algebra.tensor_product.tensor_product.is_scalar_tower) since it is automatically inferred from [tensor_product.is_scalar_tower](https://leanprover-community.github.io/mathlib_docs/linear_algebra/tensor_product.html#tensor_product.is_scalar_tower) Co-authored-by: Antoine Chambert-Loir --- src/ring_theory/tensor_product.lean | 9 ++++----- 1 file changed, 4 insertions(+), 5 deletions(-) diff --git a/src/ring_theory/tensor_product.lean b/src/ring_theory/tensor_product.lean index 75ebf694c0df5..9e3c0fdaa57d3 100644 --- a/src/ring_theory/tensor_product.lean +++ b/src/ring_theory/tensor_product.lean @@ -383,9 +383,9 @@ def include_left_ring_hom : A →+* A ⊗[R] B := map_one' := rfl, map_mul' := by simp } -variables {S : Type*} [comm_semiring S] [algebra R S] [algebra S A] [is_scalar_tower R S A] +variables {S : Type*} [comm_semiring S] [algebra S A] -instance left_algebra : algebra S (A ⊗[R] B) := +instance left_algebra [smul_comm_class R S A] : algebra S (A ⊗[R] B) := { commutes' := λ r x, begin apply tensor_product.induction_on x, @@ -408,11 +408,9 @@ instance left_algebra : algebra S (A ⊗[R] B) := instance : algebra R (A ⊗[R] B) := infer_instance @[simp] -lemma algebra_map_apply (r : S) : +lemma algebra_map_apply [smul_comm_class R S A] (r : S) : (algebra_map S (A ⊗[R] B)) r = ((algebra_map S A) r) ⊗ₜ 1 := rfl -instance : is_scalar_tower R S (A ⊗[R] B) := ⟨λ a b c, by simp⟩ - variables {C : Type v₃} [semiring C] [algebra R C] @[ext] @@ -424,6 +422,7 @@ begin simp [H], end +-- TODO: with `smul_comm_class R S A` we can have this as an `S`-algebra morphism /-- The `R`-algebra morphism `A →ₐ[R] A ⊗[R] B` sending `a` to `a ⊗ₜ 1`. -/ def include_left : A →ₐ[R] A ⊗[R] B := { commutes' := by simp, From 571e13cacbed7bf042fd3058ce27157101433842 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 2 Jun 2023 12:04:17 +0000 Subject: [PATCH 1135/1291] feat(measure_theory/measure/hausdorff): the 1-measure of a segment is its length (#18981) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Or in other words, length along a line segment is the distance between its endpoints. This also proves that the `d`-dimensional Hausdorff measure scales by a factor of `‖r‖₊ ^ d` when the set is scaled by `r` (both linearly and affinely), and that it is translation-invariant. --- src/analysis/convex/between.lean | 4 + src/measure_theory/measure/hausdorff.lean | 140 ++++++++++++++++++++++ 2 files changed, 144 insertions(+) diff --git a/src/analysis/convex/between.lean b/src/analysis/convex/between.lean index 2679b932ce622..6d7e41fd38dff 100644 --- a/src/analysis/convex/between.lean +++ b/src/analysis/convex/between.lean @@ -63,6 +63,10 @@ lemma left_mem_affine_segment (x y : P) : x ∈ affine_segment R x y := lemma right_mem_affine_segment (x y : P) : y ∈ affine_segment R x y := ⟨1, set.right_mem_Icc.2 zero_le_one, line_map_apply_one _ _⟩ +@[simp] lemma affine_segment_same (x : P) : affine_segment R x x = {x} := +by simp_rw [affine_segment, line_map_same, affine_map.coe_const, + (set.nonempty_Icc.mpr zero_le_one).image_const] + include V' variables {R} diff --git a/src/measure_theory/measure/hausdorff.lean b/src/measure_theory/measure/hausdorff.lean index 815aab7f5eb93..6b3999bc26cf6 100644 --- a/src/measure_theory/measure/hausdorff.lean +++ b/src/measure_theory/measure/hausdorff.lean @@ -3,7 +3,9 @@ Copyright (c) 2021 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ +import analysis.convex.between import measure_theory.constructions.borel_space.basic +import measure_theory.measure.haar.inner_product_space import measure_theory.measure.lebesgue.basic import topology.metric_space.holder import topology.metric_space.metric_separated @@ -384,6 +386,19 @@ begin simp only [(diam_mono hst).trans ht, le_refl, cinfi_pos] } end +lemma mk_metric_smul (m : ℝ≥0∞ → ℝ≥0∞) {c : ℝ≥0∞} (hc : c ≠ ∞) (hc' : c ≠ 0) : + (mk_metric (c • m) : outer_measure X) = c • mk_metric m := +begin + simp only [mk_metric, mk_metric', mk_metric'.pre, induced_outer_measure, + ennreal.smul_supr], + simp_rw [smul_supr, smul_bounded_by hc, smul_extend _ hc', pi.smul_apply], +end + +lemma mk_metric_nnreal_smul (m : ℝ≥0∞ → ℝ≥0∞) {c : ℝ≥0} (hc : c ≠ 0) : + (mk_metric (c • m) : outer_measure X) = c • mk_metric m := +by rw [ennreal.smul_def, ennreal.smul_def, + mk_metric_smul m (ennreal.coe_ne_top) (ennreal.coe_ne_zero.mpr hc)] + lemma isometry_map_mk_metric (m : ℝ≥0∞ → ℝ≥0∞) {f : X → Y} (hf : isometry f) (H : monotone m ∨ surjective f) : map f (mk_metric m) = restrict (range f) (mk_metric m) := @@ -770,6 +785,25 @@ lemma hausdorff_measure_image_le (h : lipschitz_with K f) {d : ℝ} (hd : 0 ≤ end lipschitz_with +open_locale pointwise + +lemma measure_theory.measure.hausdorff_measure_smul₀ + {𝕜 E : Type*} [normed_add_comm_group E] [normed_field 𝕜] [normed_space 𝕜 E] [measurable_space E] + [borel_space E] + {d : ℝ} (hd : 0 ≤ d) {r : 𝕜} (hr : r ≠ 0) (s : set E) : + μH[d] (r • s) = ‖r‖₊ ^ d • μH[d] s := +begin + suffices : ∀ {r : 𝕜}, r ≠ 0 → ∀ s : set E, μH[d] (r • s) ≤ ‖r‖₊ ^ d • μH[d] s, + { refine le_antisymm (this hr s) _, + rw [←ennreal.le_inv_smul_iff, ←nnreal.inv_rpow, ←nnnorm_inv], + { refine eq.trans_le _ (this (inv_ne_zero hr) (r • s)), + rw inv_smul_smul₀ hr }, + { simp [hr] } }, + intros r hr s, + simpa only [ennreal.smul_def, smul_eq_mul, ← ennreal.coe_rpow_of_nonneg _ hd] + using (@lipschitz_with_smul _ E _ _ _ _ r).hausdorff_measure_image_le hd s, +end + /-! ### Antilipschitz maps do not decrease Hausdorff measures and dimension -/ @@ -868,6 +902,13 @@ end isometry_equiv namespace measure_theory +@[to_additive] +lemma hausdorff_measure_smul + {α : Type*} [has_smul α X] [has_isometric_smul α X] + {d : ℝ} (c : α) (h : 0 ≤ d ∨ surjective ((•) c : X → X)) (s : set X) : + μH[d] (c • s) = μH[d] s := +(isometry_smul X c).hausdorff_measure_image h _ + /-! ### Hausdorff measure and Lebesgue measure -/ @@ -992,4 +1033,103 @@ by rw [←(volume_preserving_pi_fin_two (λ i, ℝ)).map_eq, ←(hausdorff_measure_measure_preserving_pi_fin_two (λ i, ℝ) _).map_eq, ←hausdorff_measure_pi_real, fintype.card_fin, nat.cast_two] +/-! ### Geometric results in affine spaces -/ + +section geometric +variables {𝕜 E P : Type*} + +lemma hausdorff_measure_smul_right_image [normed_add_comm_group E] [normed_space ℝ E] + [measurable_space E] [borel_space E] (v : E) (s : set ℝ) : + μH[1] ((λ r, r • v) '' s) = ‖v‖₊ • μH[1] s := +begin + obtain rfl | hv := eq_or_ne v 0, + { haveI := no_atoms_hausdorff E one_pos, + obtain rfl | hs := s.eq_empty_or_nonempty, + { simp }, + simp [hs] }, + have hn : ‖v‖ ≠ 0 := norm_ne_zero_iff.mpr hv, + -- break line_map into pieces + suffices : μH[1] ( + ((•) ‖v‖) '' (linear_map.to_span_singleton ℝ E (‖v‖⁻¹ • v) '' s)) = ‖v‖₊ • μH[1] s, + { simpa only [set.image_image, smul_comm (norm _), inv_smul_smul₀ hn, + linear_map.to_span_singleton_apply] using this }, + have iso_smul : isometry (linear_map.to_span_singleton ℝ E (‖v‖⁻¹ • v)), + { refine add_monoid_hom_class.isometry_of_norm _ (λ x, (norm_smul _ _).trans _), + rw [norm_smul, norm_inv, norm_norm, inv_mul_cancel hn, mul_one, linear_map.id_apply] }, + rw [set.image_smul, + measure.hausdorff_measure_smul₀ zero_le_one hn, nnnorm_norm, nnreal.rpow_one, + iso_smul.hausdorff_measure_image (or.inl $ zero_le_one' ℝ)], +end + +section normed_field_affine +variables [normed_field 𝕜] [normed_add_comm_group E] [normed_space 𝕜 E] [measurable_space P] +variables [metric_space P] [normed_add_torsor E P] [borel_space P] +include E + +/-- Scaling by `c` around `x` scales the measure by `‖c‖₊ ^ d`. -/ +lemma hausdorff_measure_homothety_image + {d : ℝ} (hd : 0 ≤ d) (x : P) {c : 𝕜} (hc : c ≠ 0) (s : set P) : + μH[d] (affine_map.homothety x c '' s) = ‖c‖₊ ^ d • μH[d] s := +begin + suffices : + μH[d] (isometry_equiv.vadd_const x '' (((•) c) '' ((isometry_equiv.vadd_const x).symm '' s))) + = ‖c‖₊ ^ d • μH[d] s, + { simpa only [set.image_image] }, + borelize E, + rw [isometry_equiv.hausdorff_measure_image, set.image_smul, measure.hausdorff_measure_smul₀ hd hc, + isometry_equiv.hausdorff_measure_image], +end + +lemma hausdorff_measure_homothety_preimage + {d : ℝ} (hd : 0 ≤ d) (x : P) {c : 𝕜} (hc : c ≠ 0) (s : set P) : + μH[d] (affine_map.homothety x c ⁻¹' s) = ‖c‖₊⁻¹ ^ d • μH[d] s := +begin + change μH[d] (affine_equiv.homothety_units_mul_hom x (units.mk0 c hc) ⁻¹' s) = _, + rw [←affine_equiv.image_symm, affine_equiv.coe_homothety_units_mul_hom_apply_symm, + hausdorff_measure_homothety_image hd x (_ : 𝕜ˣ).is_unit.ne_zero, units.coe_inv, units.coe_mk0, + nnnorm_inv], +end + +/-! TODO: prove `measure.map (affine_map.homothety x c) μH[d] = ‖c‖₊⁻¹ ^ d • μH[d]`, which needs a +more general version of `affine_map.homothety_continuous` -/ + +end normed_field_affine + +section real_affine +variables [normed_add_comm_group E] [normed_space ℝ E] [measurable_space P] +variables [metric_space P] [normed_add_torsor E P] [borel_space P] +include E + +/-- Mapping a set of reals along a line segment scales the measure by the length of a segment. + +This is an auxiliary result used to prove `hausdorff_measure_affine_segment`. -/ +lemma hausdorff_measure_line_map_image (x y : P) (s : set ℝ) : + μH[1] (affine_map.line_map x y '' s) = nndist x y • μH[1] s := +begin + suffices : μH[1] (isometry_equiv.vadd_const x '' ((• y -ᵥ x) '' s)) = nndist x y • μH[1] s, + { simpa only [set.image_image] }, + borelize E, + rw [isometry_equiv.hausdorff_measure_image, hausdorff_measure_smul_right_image, + nndist_eq_nnnorm_vsub' E], +end + +/-- The measure of a segment is the distance between its endpoints. -/ +@[simp] lemma hausdorff_measure_affine_segment (x y : P) : + μH[1] (affine_segment ℝ x y) = edist x y := +begin + rw [affine_segment, hausdorff_measure_line_map_image, hausdorff_measure_real, real.volume_Icc, + sub_zero, ennreal.of_real_one, ← algebra.algebra_map_eq_smul_one], + exact (edist_nndist _ _).symm, +end + +end real_affine + +/-- The measure of a segment is the distance between its endpoints. -/ +@[simp] lemma hausdorff_measure_segment {E : Type*} [normed_add_comm_group E] + [normed_space ℝ E] [measurable_space E] [borel_space E] (x y : E) : + μH[1] (segment ℝ x y) = edist x y := +by rw [←affine_segment_eq_segment, hausdorff_measure_affine_segment] + +end geometric + end measure_theory From 3d5c4a7a5fb0d982f97ed953161264f1dbd90ead Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 2 Jun 2023 14:50:10 +0000 Subject: [PATCH 1136/1291] feat(measure_theory/measure/hausdorff): invariance instances (#19145) These are a trivial consequence of the existing lemmas. The additive cases here are the ones that actually are useful, and follow from `normed_add_torsor.to_has_isometric_vadd`. --- src/measure_theory/measure/hausdorff.lean | 10 ++++++++++ 1 file changed, 10 insertions(+) diff --git a/src/measure_theory/measure/hausdorff.lean b/src/measure_theory/measure/hausdorff.lean index 6b3999bc26cf6..75c022de2443c 100644 --- a/src/measure_theory/measure/hausdorff.lean +++ b/src/measure_theory/measure/hausdorff.lean @@ -909,6 +909,16 @@ lemma hausdorff_measure_smul μH[d] (c • s) = μH[d] s := (isometry_smul X c).hausdorff_measure_image h _ +@[to_additive] +instance {d : ℝ} [group X] [has_isometric_smul X X] : is_mul_left_invariant (μH[d] : measure X) := +{ map_mul_left_eq_self := λ x, (isometry_equiv.const_smul x).map_hausdorff_measure _ } + +@[to_additive] +instance {d : ℝ} [group X] [has_isometric_smul Xᵐᵒᵖ X] : + is_mul_right_invariant (μH[d] : measure X) := +{ map_mul_right_eq_self := λ x, + (isometry_equiv.const_smul (mul_opposite.op x)).map_hausdorff_measure _ } + /-! ### Hausdorff measure and Lebesgue measure -/ From 7fdeecc0d03cd40f7a165e6cf00a4d2286db599f Mon Sep 17 00:00:00 2001 From: Riccardo Brasca Date: Fri, 2 Jun 2023 16:27:10 +0000 Subject: [PATCH 1137/1291] chore(ring_theory/root_of_unity): move and split a file (#19144) We split `ring_theory.roots_of_unity` (almost 1200 lines) into `ring_theory.roots_of_unity.basic` and `ring_theory.roots_of_unity.minpoly`. We also move `analysis.complex.roots_of_unity` to `ring_theory.roots_of_unity.complex`. --- .../100-theorems-list/16_abel_ruffini.lean | 1 + .../37_solution_of_cubic.lean | 1 - src/field_theory/abel_ruffini.lean | 2 +- .../polynomial/cyclotomic/basic.lean | 4 +- .../polynomial/cyclotomic/roots.lean | 1 + .../basic.lean} | 208 ---------------- .../roots_of_unity/complex.lean} | 2 +- src/ring_theory/roots_of_unity/minpoly.lean | 233 ++++++++++++++++++ 8 files changed, 239 insertions(+), 213 deletions(-) rename src/ring_theory/{roots_of_unity.lean => roots_of_unity/basic.lean} (79%) rename src/{analysis/complex/roots_of_unity.lean => ring_theory/roots_of_unity/complex.lean} (99%) create mode 100644 src/ring_theory/roots_of_unity/minpoly.lean diff --git a/archive/100-theorems-list/16_abel_ruffini.lean b/archive/100-theorems-list/16_abel_ruffini.lean index b377089fa8056..02b6bb4a7d250 100644 --- a/archive/100-theorems-list/16_abel_ruffini.lean +++ b/archive/100-theorems-list/16_abel_ruffini.lean @@ -6,6 +6,7 @@ Authors: Thomas Browning import analysis.calculus.local_extr import data.nat.prime_norm_num import field_theory.abel_ruffini +import ring_theory.roots_of_unity.minpoly import ring_theory.eisenstein_criterion /-! diff --git a/archive/100-theorems-list/37_solution_of_cubic.lean b/archive/100-theorems-list/37_solution_of_cubic.lean index 8c5c11284e4e6..95b58db873e73 100644 --- a/archive/100-theorems-list/37_solution_of_cubic.lean +++ b/archive/100-theorems-list/37_solution_of_cubic.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeoff Lee -/ import tactic.linear_combination -import ring_theory.roots_of_unity import ring_theory.polynomial.cyclotomic.roots /-! diff --git a/src/field_theory/abel_ruffini.lean b/src/field_theory/abel_ruffini.lean index d5d694479963d..0909b224c515d 100644 --- a/src/field_theory/abel_ruffini.lean +++ b/src/field_theory/abel_ruffini.lean @@ -6,7 +6,7 @@ Authors: Thomas Browning, Patrick Lutz import group_theory.solvable import field_theory.polynomial_galois_group -import ring_theory.roots_of_unity +import ring_theory.roots_of_unity.basic /-! # The Abel-Ruffini Theorem diff --git a/src/ring_theory/polynomial/cyclotomic/basic.lean b/src/ring_theory/polynomial/cyclotomic/basic.lean index da333193aa6e1..53f1675d35b2f 100644 --- a/src/ring_theory/polynomial/cyclotomic/basic.lean +++ b/src/ring_theory/polynomial/cyclotomic/basic.lean @@ -6,14 +6,14 @@ Authors: Riccardo Brasca import algebra.ne_zero import algebra.polynomial.big_operators -import analysis.complex.roots_of_unity +import ring_theory.roots_of_unity.complex import data.polynomial.lifts import data.polynomial.splits import data.zmod.algebra import field_theory.ratfunc import field_theory.separable import number_theory.arithmetic_function -import ring_theory.roots_of_unity +import ring_theory.roots_of_unity.basic /-! # Cyclotomic polynomials. diff --git a/src/ring_theory/polynomial/cyclotomic/roots.lean b/src/ring_theory/polynomial/cyclotomic/roots.lean index 78bdafb88dbc6..8fb945ebf74d6 100644 --- a/src/ring_theory/polynomial/cyclotomic/roots.lean +++ b/src/ring_theory/polynomial/cyclotomic/roots.lean @@ -5,6 +5,7 @@ Authors: Riccardo Brasca -/ import ring_theory.polynomial.cyclotomic.basic +import ring_theory.roots_of_unity.minpoly /-! # Roots of cyclotomic polynomials. diff --git a/src/ring_theory/roots_of_unity.lean b/src/ring_theory/roots_of_unity/basic.lean similarity index 79% rename from src/ring_theory/roots_of_unity.lean rename to src/ring_theory/roots_of_unity/basic.lean index d5fe89ae3a8ca..3c0bd5d6c363e 100644 --- a/src/ring_theory/roots_of_unity.lean +++ b/src/ring_theory/roots_of_unity/basic.lean @@ -9,7 +9,6 @@ import algebra.ne_zero import algebra.gcd_monoid.integrally_closed import data.polynomial.ring_division import field_theory.finite.basic -import field_theory.minpoly.is_integrally_closed import field_theory.separable import group_theory.specific_groups.cyclic import number_theory.divisors @@ -880,213 +879,6 @@ end end is_domain -section minpoly - -open minpoly - -section comm_ring -variables {n : ℕ} {K : Type*} [comm_ring K] {μ : K} (h : is_primitive_root μ n) (hpos : 0 < n) - -include n μ h hpos - -/--`μ` is integral over `ℤ`. -/ -lemma is_integral : is_integral ℤ μ := -begin - use (X ^ n - 1), - split, - { exact (monic_X_pow_sub_C 1 (ne_of_lt hpos).symm) }, - { simp only [((is_primitive_root.iff_def μ n).mp h).left, eval₂_one, eval₂_X_pow, eval₂_sub, - sub_self] } -end - -section is_domain - -variables [is_domain K] [char_zero K] - -omit hpos - -/--The minimal polynomial of a root of unity `μ` divides `X ^ n - 1`. -/ -lemma minpoly_dvd_X_pow_sub_one : minpoly ℤ μ ∣ X ^ n - 1 := -begin - rcases n.eq_zero_or_pos with rfl | hpos, - { simp }, - letI : is_integrally_closed ℤ := gcd_monoid.to_is_integrally_closed, - apply minpoly.is_integrally_closed_dvd (is_integral h hpos), - simp only [((is_primitive_root.iff_def μ n).mp h).left, aeval_X_pow, eq_int_cast, - int.cast_one, aeval_one, alg_hom.map_sub, sub_self] -end - -/-- The reduction modulo `p` of the minimal polynomial of a root of unity `μ` is separable. -/ -lemma separable_minpoly_mod {p : ℕ} [fact p.prime] (hdiv : ¬p ∣ n) : - separable (map (int.cast_ring_hom (zmod p)) (minpoly ℤ μ)) := -begin - have hdvd : (map (int.cast_ring_hom (zmod p)) - (minpoly ℤ μ)) ∣ X ^ n - 1, - { simpa [polynomial.map_pow, map_X, polynomial.map_one, polynomial.map_sub] using - ring_hom.map_dvd (map_ring_hom (int.cast_ring_hom (zmod p))) - (minpoly_dvd_X_pow_sub_one h) }, - refine separable.of_dvd (separable_X_pow_sub_C 1 _ one_ne_zero) hdvd, - by_contra hzero, - exact hdiv ((zmod.nat_coe_zmod_eq_zero_iff_dvd n p).1 hzero) -end - -/-- The reduction modulo `p` of the minimal polynomial of a root of unity `μ` is squarefree. -/ -lemma squarefree_minpoly_mod {p : ℕ} [fact p.prime] (hdiv : ¬ p ∣ n) : - squarefree (map (int.cast_ring_hom (zmod p)) (minpoly ℤ μ)) := -(separable_minpoly_mod h hdiv).squarefree - -/- Let `P` be the minimal polynomial of a root of unity `μ` and `Q` be the minimal polynomial of -`μ ^ p`, where `p` is a natural number that does not divide `n`. Then `P` divides `expand ℤ p Q`. -/ -lemma minpoly_dvd_expand {p : ℕ} (hdiv : ¬ p ∣ n) : minpoly ℤ μ ∣ expand ℤ p (minpoly ℤ (μ ^ p)) := -begin - rcases n.eq_zero_or_pos with rfl | hpos, - { simp * at *, }, - letI : is_integrally_closed ℤ := gcd_monoid.to_is_integrally_closed, - refine minpoly.is_integrally_closed_dvd (h.is_integral hpos) _, - { rw [aeval_def, coe_expand, ← comp, eval₂_eq_eval_map, map_comp, polynomial.map_pow, map_X, - eval_comp, eval_pow, eval_X, ← eval₂_eq_eval_map, ← aeval_def], - exact minpoly.aeval _ _ } -end - -/- Let `P` be the minimal polynomial of a root of unity `μ` and `Q` be the minimal polynomial of -`μ ^ p`, where `p` is a prime that does not divide `n`. Then `P` divides `Q ^ p` modulo `p`. -/ -lemma minpoly_dvd_pow_mod {p : ℕ} [hprime : fact p.prime] (hdiv : ¬ p ∣ n) : - map (int.cast_ring_hom (zmod p)) (minpoly ℤ μ) ∣ - map (int.cast_ring_hom (zmod p)) (minpoly ℤ (μ ^ p)) ^ p := -begin - set Q := minpoly ℤ (μ ^ p), - have hfrob : map (int.cast_ring_hom (zmod p)) Q ^ p = - map (int.cast_ring_hom (zmod p)) (expand ℤ p Q), - by rw [← zmod.expand_card, map_expand], - rw [hfrob], - apply ring_hom.map_dvd (map_ring_hom (int.cast_ring_hom (zmod p))), - exact minpoly_dvd_expand h hdiv -end - -/- Let `P` be the minimal polynomial of a root of unity `μ` and `Q` be the minimal polynomial of -`μ ^ p`, where `p` is a prime that does not divide `n`. Then `P` divides `Q` modulo `p`. -/ -lemma minpoly_dvd_mod_p {p : ℕ} [hprime : fact p.prime] (hdiv : ¬ p ∣ n) : - map (int.cast_ring_hom (zmod p)) (minpoly ℤ μ) ∣ - map (int.cast_ring_hom (zmod p)) (minpoly ℤ (μ ^ p)) := -(unique_factorization_monoid.dvd_pow_iff_dvd_of_squarefree (squarefree_minpoly_mod h - hdiv) hprime.1.ne_zero).1 (minpoly_dvd_pow_mod h hdiv) - -/-- If `p` is a prime that does not divide `n`, -then the minimal polynomials of a primitive `n`-th root of unity `μ` -and of `μ ^ p` are the same. -/ -lemma minpoly_eq_pow {p : ℕ} [hprime : fact p.prime] (hdiv : ¬ p ∣ n) : - minpoly ℤ μ = minpoly ℤ (μ ^ p) := -begin - by_cases hn : n = 0, { simp * at *, }, - have hpos := nat.pos_of_ne_zero hn, - by_contra hdiff, - set P := minpoly ℤ μ, - set Q := minpoly ℤ (μ ^ p), - have Pmonic : P.monic := minpoly.monic (h.is_integral hpos), - have Qmonic : Q.monic := minpoly.monic ((h.pow_of_prime hprime.1 hdiv).is_integral hpos), - have Pirr : irreducible P := minpoly.irreducible (h.is_integral hpos), - have Qirr : irreducible Q := - minpoly.irreducible ((h.pow_of_prime hprime.1 hdiv).is_integral hpos), - have PQprim : is_primitive (P * Q) := Pmonic.is_primitive.mul Qmonic.is_primitive, - have prod : P * Q ∣ X ^ n - 1, - { rw [(is_primitive.int.dvd_iff_map_cast_dvd_map_cast (P * Q) (X ^ n - 1) PQprim - (monic_X_pow_sub_C (1 : ℤ) (ne_of_gt hpos)).is_primitive), polynomial.map_mul], - refine is_coprime.mul_dvd _ _ _, - { have aux := is_primitive.int.irreducible_iff_irreducible_map_cast Pmonic.is_primitive, - refine (dvd_or_coprime _ _ (aux.1 Pirr)).resolve_left _, - rw map_dvd_map (int.cast_ring_hom ℚ) int.cast_injective Pmonic, - intro hdiv, - refine hdiff (eq_of_monic_of_associated Pmonic Qmonic _), - exact associated_of_dvd_dvd hdiv (Pirr.dvd_symm Qirr hdiv) }, - { apply (map_dvd_map (int.cast_ring_hom ℚ) int.cast_injective Pmonic).2, - exact minpoly_dvd_X_pow_sub_one h }, - { apply (map_dvd_map (int.cast_ring_hom ℚ) int.cast_injective Qmonic).2, - exact minpoly_dvd_X_pow_sub_one (pow_of_prime h hprime.1 hdiv) } }, - replace prod := ring_hom.map_dvd ((map_ring_hom (int.cast_ring_hom (zmod p)))) prod, - rw [coe_map_ring_hom, polynomial.map_mul, polynomial.map_sub, - polynomial.map_one, polynomial.map_pow, map_X] at prod, - obtain ⟨R, hR⟩ := minpoly_dvd_mod_p h hdiv, - rw [hR, ← mul_assoc, ← polynomial.map_mul, ← sq, polynomial.map_pow] at prod, - have habs : map (int.cast_ring_hom (zmod p)) P ^ 2 ∣ map (int.cast_ring_hom (zmod p)) P ^ 2 * R, - { use R }, - replace habs := lt_of_lt_of_le (part_enat.coe_lt_coe.2 one_lt_two) - (multiplicity.le_multiplicity_of_pow_dvd (dvd_trans habs prod)), - have hfree : squarefree (X ^ n - 1 : (zmod p)[X]), - { exact (separable_X_pow_sub_C 1 - (λ h, hdiv $ (zmod.nat_coe_zmod_eq_zero_iff_dvd n p).1 h) one_ne_zero).squarefree }, - cases (multiplicity.squarefree_iff_multiplicity_le_one (X ^ n - 1)).1 hfree - (map (int.cast_ring_hom (zmod p)) P) with hle hunit, - { rw nat.cast_one at habs, exact hle.not_lt habs }, - { replace hunit := degree_eq_zero_of_is_unit hunit, - rw degree_map_eq_of_leading_coeff_ne_zero (int.cast_ring_hom (zmod p)) _ at hunit, - { exact (minpoly.degree_pos (is_integral h hpos)).ne' hunit }, - simp only [Pmonic, eq_int_cast, monic.leading_coeff, int.cast_one, ne.def, - not_false_iff, one_ne_zero] } -end - -/-- If `m : ℕ` is coprime with `n`, -then the minimal polynomials of a primitive `n`-th root of unity `μ` -and of `μ ^ m` are the same. -/ -lemma minpoly_eq_pow_coprime {m : ℕ} (hcop : nat.coprime m n) : - minpoly ℤ μ = minpoly ℤ (μ ^ m) := -begin - revert n hcop, - refine unique_factorization_monoid.induction_on_prime m _ _ _, - { intros n hn h, - congr, - simpa [(nat.coprime_zero_left n).mp hn] using h }, - { intros u hunit n hcop h, - congr, - simp [nat.is_unit_iff.mp hunit] }, - { intros a p ha hprime hind n hcop h, - rw hind (nat.coprime.coprime_mul_left hcop) h, clear hind, - replace hprime := hprime.nat_prime, - have hdiv := (nat.prime.coprime_iff_not_dvd hprime).1 (nat.coprime.coprime_mul_right hcop), - haveI := fact.mk hprime, - rw [minpoly_eq_pow (h.pow_of_coprime a (nat.coprime.coprime_mul_left hcop)) hdiv], - congr' 1, - ring_exp } -end - -/-- If `m : ℕ` is coprime with `n`, -then the minimal polynomial of a primitive `n`-th root of unity `μ` -has `μ ^ m` as root. -/ -lemma pow_is_root_minpoly {m : ℕ} (hcop : nat.coprime m n) : - is_root (map (int.cast_ring_hom K) (minpoly ℤ μ)) (μ ^ m) := -by simpa [minpoly_eq_pow_coprime h hcop, eval_map, aeval_def (μ ^ m) _] - using minpoly.aeval ℤ (μ ^ m) - -/-- `primitive_roots n K` is a subset of the roots of the minimal polynomial of a primitive -`n`-th root of unity `μ`. -/ -lemma is_roots_of_minpoly : primitive_roots n K ⊆ (map (int.cast_ring_hom K) - (minpoly ℤ μ)).roots.to_finset := -begin - by_cases hn : n = 0, { simp * at *, }, - have hpos := nat.pos_of_ne_zero hn, - intros x hx, - obtain ⟨m, hle, hcop, rfl⟩ := (is_primitive_root_iff h hpos).1 ((mem_primitive_roots hpos).1 hx), - simpa [multiset.mem_to_finset, - mem_roots (map_monic_ne_zero $ minpoly.monic $ is_integral h hpos)] - using pow_is_root_minpoly h hcop -end - -/-- The degree of the minimal polynomial of `μ` is at least `totient n`. -/ -lemma totient_le_degree_minpoly : nat.totient n ≤ (minpoly ℤ μ).nat_degree := -let P : ℤ[X] := minpoly ℤ μ,-- minimal polynomial of `μ` - P_K : K[X] := map (int.cast_ring_hom K) P -- minimal polynomial of `μ` sent to `K[X]` -in calc -n.totient = (primitive_roots n K).card : h.card_primitive_roots.symm -... ≤ P_K.roots.to_finset.card : finset.card_le_of_subset (is_roots_of_minpoly h) -... ≤ P_K.roots.card : multiset.to_finset_card_le _ -... ≤ P_K.nat_degree : card_roots' _ -... ≤ P.nat_degree : nat_degree_map_le _ _ - -end is_domain - -end comm_ring - -end minpoly - section automorphisms variables {S} [comm_ring S] [is_domain S] {μ : S} {n : ℕ+} (hμ : is_primitive_root μ n) diff --git a/src/analysis/complex/roots_of_unity.lean b/src/ring_theory/roots_of_unity/complex.lean similarity index 99% rename from src/analysis/complex/roots_of_unity.lean rename to src/ring_theory/roots_of_unity/complex.lean index ac3d0e1e87b18..af0185b342e9c 100644 --- a/src/analysis/complex/roots_of_unity.lean +++ b/src/ring_theory/roots_of_unity/complex.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin -/ import analysis.special_functions.complex.log -import ring_theory.roots_of_unity +import ring_theory.roots_of_unity.basic /-! # Complex roots of unity diff --git a/src/ring_theory/roots_of_unity/minpoly.lean b/src/ring_theory/roots_of_unity/minpoly.lean new file mode 100644 index 0000000000000..5ab77bce66d05 --- /dev/null +++ b/src/ring_theory/roots_of_unity/minpoly.lean @@ -0,0 +1,233 @@ +/- +Copyright (c) 2020 Riccardo Brasca. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Riccardo Brasca, Johan Commelin +-/ + +import ring_theory.roots_of_unity.basic +import field_theory.minpoly.is_integrally_closed + +/-! +# Minimal polynomial of roots of unity + +We gather several results about minimal polynomial of root of unity. + +## Main results + +* `is_primitive_root.totient_le_degree_minpoly`: The degree of the minimal polynomial of a `n`-th + primitive root of unity is at least `totient n`. + +-/ + +open minpoly polynomial + +open_locale polynomial + +namespace is_primitive_root + +section comm_ring +variables {n : ℕ} {K : Type*} [comm_ring K] {μ : K} (h : is_primitive_root μ n) (hpos : 0 < n) + +include n μ h hpos + +/--`μ` is integral over `ℤ`. -/ +lemma is_integral : is_integral ℤ μ := +begin + use (X ^ n - 1), + split, + { exact (monic_X_pow_sub_C 1 (ne_of_lt hpos).symm) }, + { simp only [((is_primitive_root.iff_def μ n).mp h).left, eval₂_one, eval₂_X_pow, eval₂_sub, + sub_self] } +end + +section is_domain + +variables [is_domain K] [char_zero K] + +omit hpos + +/--The minimal polynomial of a root of unity `μ` divides `X ^ n - 1`. -/ +lemma minpoly_dvd_X_pow_sub_one : minpoly ℤ μ ∣ X ^ n - 1 := +begin + rcases n.eq_zero_or_pos with rfl | hpos, + { simp }, + letI : is_integrally_closed ℤ := gcd_monoid.to_is_integrally_closed, + apply minpoly.is_integrally_closed_dvd (is_integral h hpos), + simp only [((is_primitive_root.iff_def μ n).mp h).left, aeval_X_pow, eq_int_cast, + int.cast_one, aeval_one, alg_hom.map_sub, sub_self] +end + +/-- The reduction modulo `p` of the minimal polynomial of a root of unity `μ` is separable. -/ +lemma separable_minpoly_mod {p : ℕ} [fact p.prime] (hdiv : ¬p ∣ n) : + separable (map (int.cast_ring_hom (zmod p)) (minpoly ℤ μ)) := +begin + have hdvd : (map (int.cast_ring_hom (zmod p)) + (minpoly ℤ μ)) ∣ X ^ n - 1, + { simpa [polynomial.map_pow, map_X, polynomial.map_one, polynomial.map_sub] using + ring_hom.map_dvd (map_ring_hom (int.cast_ring_hom (zmod p))) + (minpoly_dvd_X_pow_sub_one h) }, + refine separable.of_dvd (separable_X_pow_sub_C 1 _ one_ne_zero) hdvd, + by_contra hzero, + exact hdiv ((zmod.nat_coe_zmod_eq_zero_iff_dvd n p).1 hzero) +end + +/-- The reduction modulo `p` of the minimal polynomial of a root of unity `μ` is squarefree. -/ +lemma squarefree_minpoly_mod {p : ℕ} [fact p.prime] (hdiv : ¬ p ∣ n) : + squarefree (map (int.cast_ring_hom (zmod p)) (minpoly ℤ μ)) := +(separable_minpoly_mod h hdiv).squarefree + +/- Let `P` be the minimal polynomial of a root of unity `μ` and `Q` be the minimal polynomial of +`μ ^ p`, where `p` is a natural number that does not divide `n`. Then `P` divides `expand ℤ p Q`. -/ +lemma minpoly_dvd_expand {p : ℕ} (hdiv : ¬ p ∣ n) : minpoly ℤ μ ∣ expand ℤ p (minpoly ℤ (μ ^ p)) := +begin + rcases n.eq_zero_or_pos with rfl | hpos, + { simp * at *, }, + letI : is_integrally_closed ℤ := gcd_monoid.to_is_integrally_closed, + refine minpoly.is_integrally_closed_dvd (h.is_integral hpos) _, + { rw [aeval_def, coe_expand, ← comp, eval₂_eq_eval_map, map_comp, polynomial.map_pow, map_X, + eval_comp, eval_pow, eval_X, ← eval₂_eq_eval_map, ← aeval_def], + exact minpoly.aeval _ _ } +end + +/- Let `P` be the minimal polynomial of a root of unity `μ` and `Q` be the minimal polynomial of +`μ ^ p`, where `p` is a prime that does not divide `n`. Then `P` divides `Q ^ p` modulo `p`. -/ +lemma minpoly_dvd_pow_mod {p : ℕ} [hprime : fact p.prime] (hdiv : ¬ p ∣ n) : + map (int.cast_ring_hom (zmod p)) (minpoly ℤ μ) ∣ + map (int.cast_ring_hom (zmod p)) (minpoly ℤ (μ ^ p)) ^ p := +begin + set Q := minpoly ℤ (μ ^ p), + have hfrob : map (int.cast_ring_hom (zmod p)) Q ^ p = + map (int.cast_ring_hom (zmod p)) (expand ℤ p Q), + by rw [← zmod.expand_card, map_expand], + rw [hfrob], + apply ring_hom.map_dvd (map_ring_hom (int.cast_ring_hom (zmod p))), + exact minpoly_dvd_expand h hdiv +end + +/- Let `P` be the minimal polynomial of a root of unity `μ` and `Q` be the minimal polynomial of +`μ ^ p`, where `p` is a prime that does not divide `n`. Then `P` divides `Q` modulo `p`. -/ +lemma minpoly_dvd_mod_p {p : ℕ} [hprime : fact p.prime] (hdiv : ¬ p ∣ n) : + map (int.cast_ring_hom (zmod p)) (minpoly ℤ μ) ∣ + map (int.cast_ring_hom (zmod p)) (minpoly ℤ (μ ^ p)) := +(unique_factorization_monoid.dvd_pow_iff_dvd_of_squarefree (squarefree_minpoly_mod h + hdiv) hprime.1.ne_zero).1 (minpoly_dvd_pow_mod h hdiv) + +/-- If `p` is a prime that does not divide `n`, +then the minimal polynomials of a primitive `n`-th root of unity `μ` +and of `μ ^ p` are the same. -/ +lemma minpoly_eq_pow {p : ℕ} [hprime : fact p.prime] (hdiv : ¬ p ∣ n) : + minpoly ℤ μ = minpoly ℤ (μ ^ p) := +begin + classical, + by_cases hn : n = 0, { simp * at *, }, + have hpos := nat.pos_of_ne_zero hn, + by_contra hdiff, + set P := minpoly ℤ μ, + set Q := minpoly ℤ (μ ^ p), + have Pmonic : P.monic := minpoly.monic (h.is_integral hpos), + have Qmonic : Q.monic := minpoly.monic ((h.pow_of_prime hprime.1 hdiv).is_integral hpos), + have Pirr : irreducible P := minpoly.irreducible (h.is_integral hpos), + have Qirr : irreducible Q := + minpoly.irreducible ((h.pow_of_prime hprime.1 hdiv).is_integral hpos), + have PQprim : is_primitive (P * Q) := Pmonic.is_primitive.mul Qmonic.is_primitive, + have prod : P * Q ∣ X ^ n - 1, + { rw [(is_primitive.int.dvd_iff_map_cast_dvd_map_cast (P * Q) (X ^ n - 1) PQprim + (monic_X_pow_sub_C (1 : ℤ) (ne_of_gt hpos)).is_primitive), polynomial.map_mul], + refine is_coprime.mul_dvd _ _ _, + { have aux := is_primitive.int.irreducible_iff_irreducible_map_cast Pmonic.is_primitive, + refine (dvd_or_coprime _ _ (aux.1 Pirr)).resolve_left _, + rw map_dvd_map (int.cast_ring_hom ℚ) int.cast_injective Pmonic, + intro hdiv, + refine hdiff (eq_of_monic_of_associated Pmonic Qmonic _), + exact associated_of_dvd_dvd hdiv (Pirr.dvd_symm Qirr hdiv) }, + { apply (map_dvd_map (int.cast_ring_hom ℚ) int.cast_injective Pmonic).2, + exact minpoly_dvd_X_pow_sub_one h }, + { apply (map_dvd_map (int.cast_ring_hom ℚ) int.cast_injective Qmonic).2, + exact minpoly_dvd_X_pow_sub_one (pow_of_prime h hprime.1 hdiv) } }, + replace prod := ring_hom.map_dvd ((map_ring_hom (int.cast_ring_hom (zmod p)))) prod, + rw [coe_map_ring_hom, polynomial.map_mul, polynomial.map_sub, + polynomial.map_one, polynomial.map_pow, map_X] at prod, + obtain ⟨R, hR⟩ := minpoly_dvd_mod_p h hdiv, + rw [hR, ← mul_assoc, ← polynomial.map_mul, ← sq, polynomial.map_pow] at prod, + have habs : map (int.cast_ring_hom (zmod p)) P ^ 2 ∣ map (int.cast_ring_hom (zmod p)) P ^ 2 * R, + { use R }, + replace habs := lt_of_lt_of_le (part_enat.coe_lt_coe.2 one_lt_two) + (multiplicity.le_multiplicity_of_pow_dvd (dvd_trans habs prod)), + have hfree : squarefree (X ^ n - 1 : (zmod p)[X]), + { exact (separable_X_pow_sub_C 1 + (λ h, hdiv $ (zmod.nat_coe_zmod_eq_zero_iff_dvd n p).1 h) one_ne_zero).squarefree }, + cases (multiplicity.squarefree_iff_multiplicity_le_one (X ^ n - 1)).1 hfree + (map (int.cast_ring_hom (zmod p)) P) with hle hunit, + { rw nat.cast_one at habs, exact hle.not_lt habs }, + { replace hunit := degree_eq_zero_of_is_unit hunit, + rw degree_map_eq_of_leading_coeff_ne_zero (int.cast_ring_hom (zmod p)) _ at hunit, + { exact (minpoly.degree_pos (is_integral h hpos)).ne' hunit }, + simp only [Pmonic, eq_int_cast, monic.leading_coeff, int.cast_one, ne.def, + not_false_iff, one_ne_zero] } +end + +/-- If `m : ℕ` is coprime with `n`, +then the minimal polynomials of a primitive `n`-th root of unity `μ` +and of `μ ^ m` are the same. -/ +lemma minpoly_eq_pow_coprime {m : ℕ} (hcop : nat.coprime m n) : + minpoly ℤ μ = minpoly ℤ (μ ^ m) := +begin + revert n hcop, + refine unique_factorization_monoid.induction_on_prime m _ _ _, + { intros n hn h, + congr, + simpa [(nat.coprime_zero_left n).mp hn] using h }, + { intros u hunit n hcop h, + congr, + simp [nat.is_unit_iff.mp hunit] }, + { intros a p ha hprime hind n hcop h, + rw hind (nat.coprime.coprime_mul_left hcop) h, clear hind, + replace hprime := hprime.nat_prime, + have hdiv := (nat.prime.coprime_iff_not_dvd hprime).1 (nat.coprime.coprime_mul_right hcop), + haveI := fact.mk hprime, + rw [minpoly_eq_pow (h.pow_of_coprime a (nat.coprime.coprime_mul_left hcop)) hdiv], + congr' 1, + ring_exp } +end + +/-- If `m : ℕ` is coprime with `n`, +then the minimal polynomial of a primitive `n`-th root of unity `μ` +has `μ ^ m` as root. -/ +lemma pow_is_root_minpoly {m : ℕ} (hcop : nat.coprime m n) : + is_root (map (int.cast_ring_hom K) (minpoly ℤ μ)) (μ ^ m) := +by simpa [minpoly_eq_pow_coprime h hcop, eval_map, aeval_def (μ ^ m) _] + using minpoly.aeval ℤ (μ ^ m) + +/-- `primitive_roots n K` is a subset of the roots of the minimal polynomial of a primitive +`n`-th root of unity `μ`. -/ +lemma is_roots_of_minpoly [decidable_eq K] : primitive_roots n K ⊆ (map (int.cast_ring_hom K) + (minpoly ℤ μ)).roots.to_finset := +begin + by_cases hn : n = 0, { simp * at *, }, + have hpos := nat.pos_of_ne_zero hn, + intros x hx, + obtain ⟨m, hle, hcop, rfl⟩ := (is_primitive_root_iff h hpos).1 ((mem_primitive_roots hpos).1 hx), + simpa [multiset.mem_to_finset, + mem_roots (map_monic_ne_zero $ minpoly.monic $ is_integral h hpos)] + using pow_is_root_minpoly h hcop +end + +/-- The degree of the minimal polynomial of `μ` is at least `totient n`. -/ +lemma totient_le_degree_minpoly : nat.totient n ≤ (minpoly ℤ μ).nat_degree := +begin + classical, + let P : ℤ[X] := minpoly ℤ μ,-- minimal polynomial of `μ` + let P_K : K[X] := map (int.cast_ring_hom K) P, -- minimal polynomial of `μ` sent to `K[X]` + calc + n.totient = (primitive_roots n K).card : h.card_primitive_roots.symm + ... ≤ P_K.roots.to_finset.card : finset.card_le_of_subset (is_roots_of_minpoly h) + ... ≤ P_K.roots.card : multiset.to_finset_card_le _ + ... ≤ P_K.nat_degree : card_roots' _ + ... ≤ P.nat_degree : nat_degree_map_le _ _ +end + +end is_domain + +end comm_ring + +end is_primitive_root From af471b9e3ce868f296626d33189b4ce730fa4c00 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Sat, 3 Jun 2023 13:39:51 +0000 Subject: [PATCH 1138/1291] chore(*): add mathlib4 synchronization comments (#19148) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.gcd_monoid.integrally_closed` * `algebra.lie.submodule` * `algebraic_geometry.locally_ringed_space` * `algebraic_geometry.ringed_space` * `analysis.analytic.linear` * `analysis.analytic.radius_liminf` * `analysis.calculus.extend_deriv` * `analysis.calculus.lhopital` * `analysis.calculus.taylor` * `analysis.calculus.uniform_limits_deriv` * `analysis.inner_product_space.l2_space` * `analysis.inner_product_space.orientation` * `analysis.normed_space.exponential` * `analysis.normed_space.star.exponential` * `analysis.von_neumann_algebra.basic` * `category_theory.extensive` * `category_theory.monoidal.braided` * `category_theory.monoidal.rigid.basic` * `data.mv_polynomial.derivation` * `data.ordmap.ordset` * `field_theory.finite.basic` * `field_theory.finite.polynomial` * `linear_algebra.matrix.charpoly.finite_field` * `measure_theory.function.simple_func_dense_lp` * `measure_theory.function.strongly_measurable.lp` * `measure_theory.function.uniform_integrable` * `measure_theory.integral.set_to_l1` * `measure_theory.measure.lebesgue.basic` * `measure_theory.measure.lebesgue.complex` * `number_theory.liouville.basic` * `number_theory.liouville.liouville_with` * `number_theory.liouville.residual` * `probability.independence.basic` * `probability.independence.zero_one` * `ring_theory.derivation.basic` * `ring_theory.integrally_closed` * `ring_theory.polynomial.rational_root` * `ring_theory.valuation.integral` * `topology.homotopy.H_spaces` --- src/algebra/gcd_monoid/integrally_closed.lean | 3 +++ src/algebra/lie/submodule.lean | 3 +++ src/algebraic_geometry/locally_ringed_space.lean | 3 +++ src/algebraic_geometry/ringed_space.lean | 3 +++ src/analysis/analytic/linear.lean | 3 +++ src/analysis/analytic/radius_liminf.lean | 3 +++ src/analysis/calculus/extend_deriv.lean | 3 +++ src/analysis/calculus/lhopital.lean | 3 +++ src/analysis/calculus/taylor.lean | 3 +++ src/analysis/calculus/uniform_limits_deriv.lean | 3 +++ src/analysis/inner_product_space/l2_space.lean | 3 +++ src/analysis/inner_product_space/orientation.lean | 3 +++ src/analysis/normed_space/exponential.lean | 3 +++ src/analysis/normed_space/star/exponential.lean | 3 +++ src/analysis/von_neumann_algebra/basic.lean | 3 +++ src/category_theory/extensive.lean | 3 +++ src/category_theory/monoidal/braided.lean | 3 +++ src/category_theory/monoidal/rigid/basic.lean | 3 +++ src/data/mv_polynomial/derivation.lean | 3 +++ src/data/ordmap/ordset.lean | 3 +++ src/field_theory/finite/basic.lean | 3 +++ src/field_theory/finite/polynomial.lean | 3 +++ src/linear_algebra/matrix/charpoly/finite_field.lean | 3 +++ src/measure_theory/function/simple_func_dense_lp.lean | 3 +++ src/measure_theory/function/strongly_measurable/lp.lean | 3 +++ src/measure_theory/function/uniform_integrable.lean | 3 +++ src/measure_theory/integral/set_to_l1.lean | 3 +++ src/measure_theory/measure/lebesgue/basic.lean | 3 +++ src/measure_theory/measure/lebesgue/complex.lean | 3 +++ src/number_theory/liouville/basic.lean | 3 +++ src/number_theory/liouville/liouville_with.lean | 3 +++ src/number_theory/liouville/residual.lean | 3 +++ src/probability/independence/basic.lean | 3 +++ src/probability/independence/zero_one.lean | 3 +++ src/ring_theory/derivation/basic.lean | 3 +++ src/ring_theory/integrally_closed.lean | 3 +++ src/ring_theory/polynomial/rational_root.lean | 3 +++ src/ring_theory/valuation/integral.lean | 3 +++ src/topology/homotopy/H_spaces.lean | 3 +++ 39 files changed, 117 insertions(+) diff --git a/src/algebra/gcd_monoid/integrally_closed.lean b/src/algebra/gcd_monoid/integrally_closed.lean index 47ff732f99523..ac3990150f937 100644 --- a/src/algebra/gcd_monoid/integrally_closed.lean +++ b/src/algebra/gcd_monoid/integrally_closed.lean @@ -11,6 +11,9 @@ import ring_theory.polynomial.eisenstein.basic # GCD domains are integrally closed +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + -/ open_locale big_operators polynomial diff --git a/src/algebra/lie/submodule.lean b/src/algebra/lie/submodule.lean index 731c560165180..cf38862aee369 100644 --- a/src/algebra/lie/submodule.lean +++ b/src/algebra/lie/submodule.lean @@ -9,6 +9,9 @@ import ring_theory.noetherian /-! # Lie submodules of a Lie algebra +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define Lie submodules and Lie ideals, we construct the lattice structure on Lie submodules and we use it to define various important operations, notably the Lie span of a subset of a Lie module. diff --git a/src/algebraic_geometry/locally_ringed_space.lean b/src/algebraic_geometry/locally_ringed_space.lean index 3e88bfb7a6cdf..65c77096f56c8 100644 --- a/src/algebraic_geometry/locally_ringed_space.lean +++ b/src/algebraic_geometry/locally_ringed_space.lean @@ -10,6 +10,9 @@ import algebraic_geometry.stalks /-! # The category of locally ringed spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define (bundled) locally ringed spaces (as `SheafedSpace CommRing` along with the fact that the stalks are local rings), and morphisms between these (morphisms in `SheafedSpace` with `is_local_ring_hom` on the stalk maps). diff --git a/src/algebraic_geometry/ringed_space.lean b/src/algebraic_geometry/ringed_space.lean index 9018859254f69..ff9354559ca56 100644 --- a/src/algebraic_geometry/ringed_space.lean +++ b/src/algebraic_geometry/ringed_space.lean @@ -12,6 +12,9 @@ import algebra.category.Ring.limits /-! # Ringed spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We introduce the category of ringed spaces, as an alias for `SheafedSpace CommRing`. The facts collected in this file are typically stated for locally ringed spaces, but never actually diff --git a/src/analysis/analytic/linear.lean b/src/analysis/analytic/linear.lean index e0c124303f844..978b53cba4ba2 100644 --- a/src/analysis/analytic/linear.lean +++ b/src/analysis/analytic/linear.lean @@ -8,6 +8,9 @@ import analysis.analytic.basic /-! # Linear functions are analytic +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove that a `continuous_linear_map` defines an analytic function with the formal power series `f x = f a + f (x - a)`. -/ diff --git a/src/analysis/analytic/radius_liminf.lean b/src/analysis/analytic/radius_liminf.lean index 949ad9c17f72e..bbaad072c93de 100644 --- a/src/analysis/analytic/radius_liminf.lean +++ b/src/analysis/analytic/radius_liminf.lean @@ -9,6 +9,9 @@ import analysis.special_functions.pow.nnreal /-! # Representation of `formal_multilinear_series.radius` as a `liminf` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove that the radius of convergence of a `formal_multilinear_series` is equal to $\liminf_{n\to\infty} \frac{1}{\sqrt[n]{‖p n‖}}$. This lemma can't go to `basic.lean` because this would create a circular dependency once we redefine `exp` using `formal_multilinear_series`. diff --git a/src/analysis/calculus/extend_deriv.lean b/src/analysis/calculus/extend_deriv.lean index a80b6895a99a2..b2c01107784ca 100644 --- a/src/analysis/calculus/extend_deriv.lean +++ b/src/analysis/calculus/extend_deriv.lean @@ -8,6 +8,9 @@ import analysis.calculus.mean_value /-! # Extending differentiability to the boundary +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We investigate how differentiable functions inside a set extend to differentiable functions on the boundary. For this, it suffices that the function and its derivative admit limits there. A general version of this statement is given in `has_fderiv_at_boundary_of_tendsto_fderiv`. diff --git a/src/analysis/calculus/lhopital.lean b/src/analysis/calculus/lhopital.lean index 5c2b6d617e122..8ab77bf182d86 100644 --- a/src/analysis/calculus/lhopital.lean +++ b/src/analysis/calculus/lhopital.lean @@ -9,6 +9,9 @@ import analysis.calculus.deriv.inv /-! # L'Hôpital's rule for 0/0 indeterminate forms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we prove several forms of "L'Hopital's rule" for computing 0/0 indeterminate forms. The proof of `has_deriv_at.lhopital_zero_right_on_Ioo` is based on the one given in the corresponding diff --git a/src/analysis/calculus/taylor.lean b/src/analysis/calculus/taylor.lean index eb40d84acfff6..f7040ff75e28c 100644 --- a/src/analysis/calculus/taylor.lean +++ b/src/analysis/calculus/taylor.lean @@ -10,6 +10,9 @@ import data.polynomial.module /-! # Taylor's theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the Taylor polynomial of a real function `f : ℝ → E`, where `E` is a normed vector space over `ℝ` and proves Taylor's theorem, which states that if `f` is sufficiently smooth, then diff --git a/src/analysis/calculus/uniform_limits_deriv.lean b/src/analysis/calculus/uniform_limits_deriv.lean index cf8881f0e5a3e..a5b1a1b9b99a3 100644 --- a/src/analysis/calculus/uniform_limits_deriv.lean +++ b/src/analysis/calculus/uniform_limits_deriv.lean @@ -10,6 +10,9 @@ import order.filter.curry /-! # Swapping limits and derivatives via uniform convergence +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The purpose of this file is to prove that the derivative of the pointwise limit of a sequence of functions is the pointwise limit of the functions' derivatives when the derivatives converge _uniformly_. The formal statement appears as `has_fderiv_at_of_tendsto_locally_uniformly_at`. diff --git a/src/analysis/inner_product_space/l2_space.lean b/src/analysis/inner_product_space/l2_space.lean index 1dba544d0c830..10d087817b0ce 100644 --- a/src/analysis/inner_product_space/l2_space.lean +++ b/src/analysis/inner_product_space/l2_space.lean @@ -10,6 +10,9 @@ import analysis.inner_product_space.pi_L2 /-! # Hilbert sum of a family of inner product spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given a family `(G : ι → Type*) [Π i, inner_product_space 𝕜 (G i)]` of inner product spaces, this file equips `lp G 2` with an inner product space structure, where `lp G 2` consists of those dependent functions `f : Π i, G i` for which `∑' i, ‖f i‖ ^ 2`, the sum of the norms-squared, is diff --git a/src/analysis/inner_product_space/orientation.lean b/src/analysis/inner_product_space/orientation.lean index 7066e98cf60d5..db8a8457ae7c7 100644 --- a/src/analysis/inner_product_space/orientation.lean +++ b/src/analysis/inner_product_space/orientation.lean @@ -9,6 +9,9 @@ import linear_algebra.orientation /-! # Orientations of real inner product spaces. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file provides definitions and proves lemmas about orientations of real inner product spaces. ## Main definitions diff --git a/src/analysis/normed_space/exponential.lean b/src/analysis/normed_space/exponential.lean index b0d62cf8427b4..f51e1c0c4a303 100644 --- a/src/analysis/normed_space/exponential.lean +++ b/src/analysis/normed_space/exponential.lean @@ -13,6 +13,9 @@ import topology.algebra.algebra /-! # Exponential in a Banach algebra +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we define `exp 𝕂 : 𝔸 → 𝔸`, the exponential map in a topological algebra `𝔸` over a field `𝕂`. diff --git a/src/analysis/normed_space/star/exponential.lean b/src/analysis/normed_space/star/exponential.lean index 36b05dd9cf7b1..1c07409b88ab5 100644 --- a/src/analysis/normed_space/star/exponential.lean +++ b/src/analysis/normed_space/star/exponential.lean @@ -6,6 +6,9 @@ Authors: Jireh Loreaux import analysis.normed_space.exponential /-! # The exponential map from selfadjoint to unitary + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. In this file, we establish various propreties related to the map `λ a, exp ℂ A (I • a)` between the subtypes `self_adjoint A` and `unitary A`. diff --git a/src/analysis/von_neumann_algebra/basic.lean b/src/analysis/von_neumann_algebra/basic.lean index 84dc25e66e8be..dc3ec7acc5250 100644 --- a/src/analysis/von_neumann_algebra/basic.lean +++ b/src/analysis/von_neumann_algebra/basic.lean @@ -12,6 +12,9 @@ import algebra.star.subalgebra /-! # Von Neumann algebras +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We give the "abstract" and "concrete" definitions of a von Neumann algebra. We still have a major project ahead of us to show the equivalence between these definitions! diff --git a/src/category_theory/extensive.lean b/src/category_theory/extensive.lean index 6623631672aed..9a7349ba8495f 100644 --- a/src/category_theory/extensive.lean +++ b/src/category_theory/extensive.lean @@ -13,6 +13,9 @@ import category_theory.limits.functor_category # Extensive categories +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions - `category_theory.is_van_kampen_colimit`: A (colimit) cocone over a diagram `F : J ⥤ C` is van Kampen if for every cocone `c'` over the pullback of the diagram `F' : J ⥤ C'`, diff --git a/src/category_theory/monoidal/braided.lean b/src/category_theory/monoidal/braided.lean index 79d3d2935f8c2..bc15bf64ccc39 100644 --- a/src/category_theory/monoidal/braided.lean +++ b/src/category_theory/monoidal/braided.lean @@ -10,6 +10,9 @@ import category_theory.monoidal.discrete /-! # Braided and symmetric monoidal categories +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The basic definitions of braided monoidal categories, and symmetric monoidal categories, as well as braided functors. diff --git a/src/category_theory/monoidal/rigid/basic.lean b/src/category_theory/monoidal/rigid/basic.lean index a8c467d0f0362..60a4001ff3ccc 100644 --- a/src/category_theory/monoidal/rigid/basic.lean +++ b/src/category_theory/monoidal/rigid/basic.lean @@ -10,6 +10,9 @@ import tactic.apply_fun /-! # Rigid (autonomous) monoidal categories +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines rigid (autonomous) monoidal categories and the necessary theory about exact pairings and duals. diff --git a/src/data/mv_polynomial/derivation.lean b/src/data/mv_polynomial/derivation.lean index 0591c3791fb13..8302eef772b37 100644 --- a/src/data/mv_polynomial/derivation.lean +++ b/src/data/mv_polynomial/derivation.lean @@ -9,6 +9,9 @@ import ring_theory.derivation.basic /-! # Derivations of multivariate polynomials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove that a derivation of `mv_polynomial σ R` is determined by its values on all monomials `mv_polynomial.X i`. We also provide a constructor `mv_polynomial.mk_derivation` that builds a derivation from its values on `X i`s and a linear equivalence diff --git a/src/data/ordmap/ordset.lean b/src/data/ordmap/ordset.lean index bf6dbbdba64a5..c466dcbbc1b55 100644 --- a/src/data/ordmap/ordset.lean +++ b/src/data/ordmap/ordset.lean @@ -11,6 +11,9 @@ import tactic.linarith /-! # Verification of the `ordnode α` datatype +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves the correctness of the operations in `data.ordmap.ordnode`. The public facing version is the type `ordset α`, which is a wrapper around `ordnode α` which includes the correctness invariant of the type, and it exposes diff --git a/src/field_theory/finite/basic.lean b/src/field_theory/finite/basic.lean index 75b448e404c72..505d9e97e807c 100644 --- a/src/field_theory/finite/basic.lean +++ b/src/field_theory/finite/basic.lean @@ -10,6 +10,9 @@ import tactic.apply_fun /-! # Finite fields +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains basic results about finite fields. Throughout most of this file, `K` denotes a finite field and `q` is notation for the cardinality of `K`. diff --git a/src/field_theory/finite/polynomial.lean b/src/field_theory/finite/polynomial.lean index 97d574232fdc3..5576616d7209b 100644 --- a/src/field_theory/finite/polynomial.lean +++ b/src/field_theory/finite/polynomial.lean @@ -12,6 +12,9 @@ import field_theory.finite.basic /-! ## Polynomials over finite fields + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ namespace mv_polynomial diff --git a/src/linear_algebra/matrix/charpoly/finite_field.lean b/src/linear_algebra/matrix/charpoly/finite_field.lean index 010277b4b321c..429b18e1a6f64 100644 --- a/src/linear_algebra/matrix/charpoly/finite_field.lean +++ b/src/linear_algebra/matrix/charpoly/finite_field.lean @@ -10,6 +10,9 @@ import data.matrix.char_p /-! # Results on characteristic polynomials and traces over finite fields. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ noncomputable theory diff --git a/src/measure_theory/function/simple_func_dense_lp.lean b/src/measure_theory/function/simple_func_dense_lp.lean index 0c9b2cf2b6b9a..e67f84473c68a 100644 --- a/src/measure_theory/function/simple_func_dense_lp.lean +++ b/src/measure_theory/function/simple_func_dense_lp.lean @@ -9,6 +9,9 @@ import measure_theory.function.simple_func_dense /-! # Density of simple functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Show that each `Lᵖ` Borel measurable function can be approximated in `Lᵖ` norm by a sequence of simple functions. diff --git a/src/measure_theory/function/strongly_measurable/lp.lean b/src/measure_theory/function/strongly_measurable/lp.lean index d6cea24301e9e..fbcd743140127 100644 --- a/src/measure_theory/function/strongly_measurable/lp.lean +++ b/src/measure_theory/function/strongly_measurable/lp.lean @@ -10,6 +10,9 @@ import measure_theory.function.strongly_measurable.basic /-! # Finitely strongly measurable functions in `Lp` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Functions in `Lp` for `0 < p < ∞` are finitely strongly measurable. ## Main statements diff --git a/src/measure_theory/function/uniform_integrable.lean b/src/measure_theory/function/uniform_integrable.lean index 6087380fdc8a0..27cf1d4f84a64 100644 --- a/src/measure_theory/function/uniform_integrable.lean +++ b/src/measure_theory/function/uniform_integrable.lean @@ -9,6 +9,9 @@ import measure_theory.function.l1_space /-! # Uniform integrability +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains the definitions for uniform integrability (both in the measure theory sense as well as the probability theory sense). This file also contains the Vitali convergence theorem which estabishes a relation between uniform integrability, convergence in measure and diff --git a/src/measure_theory/integral/set_to_l1.lean b/src/measure_theory/integral/set_to_l1.lean index 4b2fece04f2c5..fdb5dcd6ab6d7 100644 --- a/src/measure_theory/integral/set_to_l1.lean +++ b/src/measure_theory/integral/set_to_l1.lean @@ -8,6 +8,9 @@ import measure_theory.function.simple_func_dense_lp /-! # Extension of a linear function from indicators to L1 +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Let `T : set α → E →L[ℝ] F` be additive for measurable sets with finite measure, in the sense that for `s, t` two such sets, `s ∩ t = ∅ → T (s ∪ t) = T s + T t`. `T` is akin to a bilinear map on `set α × E`, or a linear map on indicator functions. diff --git a/src/measure_theory/measure/lebesgue/basic.lean b/src/measure_theory/measure/lebesgue/basic.lean index 01c2ff63e5f03..ba416fc42a780 100644 --- a/src/measure_theory/measure/lebesgue/basic.lean +++ b/src/measure_theory/measure/lebesgue/basic.lean @@ -14,6 +14,9 @@ import measure_theory.measure.haar.of_basis /-! # Lebesgue measure on the real line and on `ℝⁿ` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We show that the Lebesgue measure on the real line (constructed as a particular case of additive Haar measure on inner product spaces) coincides with the Stieltjes measure associated to the function `x ↦ x`. We deduce properties of this measure on `ℝ`, and then of the product diff --git a/src/measure_theory/measure/lebesgue/complex.lean b/src/measure_theory/measure/lebesgue/complex.lean index 7f201d245816b..bc7a45288ab0a 100644 --- a/src/measure_theory/measure/lebesgue/complex.lean +++ b/src/measure_theory/measure/lebesgue/complex.lean @@ -10,6 +10,9 @@ import measure_theory.measure.haar.of_basis /-! # Lebesgue measure on `ℂ` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define Lebesgue measure on `ℂ`. Since `ℂ` is defined as a `structure` as the push-forward of the volume on `ℝ²` under the natural isomorphism. There are (at least) two frequently used ways to represent `ℝ²` in `mathlib`: `ℝ × ℝ` and `fin 2 → ℝ`. We define measurable diff --git a/src/number_theory/liouville/basic.lean b/src/number_theory/liouville/basic.lean index 438cb18b980fb..2f52313ceaa08 100644 --- a/src/number_theory/liouville/basic.lean +++ b/src/number_theory/liouville/basic.lean @@ -10,6 +10,9 @@ import data.real.irrational # Liouville's theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains a proof of Liouville's theorem stating that all Liouville numbers are transcendental. diff --git a/src/number_theory/liouville/liouville_with.lean b/src/number_theory/liouville/liouville_with.lean index a0f50ae4c916f..f202e4554a99c 100644 --- a/src/number_theory/liouville/liouville_with.lean +++ b/src/number_theory/liouville/liouville_with.lean @@ -10,6 +10,9 @@ import topology.instances.irrational /-! # Liouville numbers with a given exponent +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We say that a real number `x` is a Liouville number with exponent `p : ℝ` if there exists a real number `C` such that for infinitely many denominators `n` there exists a numerator `m` such that `x ≠ m / n` and `|x - m / n| < C / n ^ p`. A number is a Liouville number in the sense of diff --git a/src/number_theory/liouville/residual.lean b/src/number_theory/liouville/residual.lean index a00e4c72587c3..7188ad925b896 100644 --- a/src/number_theory/liouville/residual.lean +++ b/src/number_theory/liouville/residual.lean @@ -10,6 +10,9 @@ import topology.instances.irrational /-! # Density of Liouville numbers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove that the set of Liouville numbers form a dense `Gδ` set. We also prove a similar statement about irrational numbers. -/ diff --git a/src/probability/independence/basic.lean b/src/probability/independence/basic.lean index d7827c9312e8b..7223f9b23c01e 100644 --- a/src/probability/independence/basic.lean +++ b/src/probability/independence/basic.lean @@ -8,6 +8,9 @@ import measure_theory.constructions.pi /-! # Independence of sets of sets and measure spaces (σ-algebras) +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + * A family of sets of sets `π : ι → set (set Ω)` is independent with respect to a measure `μ` if for any finite set of indices `s = {i_1, ..., i_n}`, for any sets `f i_1 ∈ π i_1, ..., f i_n ∈ π i_n`, `μ (⋂ i in s, f i) = ∏ i in s, μ (f i) `. It will be used for families of π-systems. diff --git a/src/probability/independence/zero_one.lean b/src/probability/independence/zero_one.lean index a806d92324928..10fe4002eda6d 100644 --- a/src/probability/independence/zero_one.lean +++ b/src/probability/independence/zero_one.lean @@ -8,6 +8,9 @@ import probability.independence.basic /-! # Kolmogorov's 0-1 law +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Let `s : ι → measurable_space Ω` be an independent sequence of sub-σ-algebras. Then any set which is measurable with respect to the tail σ-algebra `limsup s at_top` has probability 0 or 1. diff --git a/src/ring_theory/derivation/basic.lean b/src/ring_theory/derivation/basic.lean index 09af046a4d1b2..95a1d1ed3e837 100644 --- a/src/ring_theory/derivation/basic.lean +++ b/src/ring_theory/derivation/basic.lean @@ -9,6 +9,9 @@ import ring_theory.adjoin.basic /-! # Derivations +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines derivation. A derivation `D` from the `R`-algebra `A` to the `A`-module `M` is an `R`-linear map that satisfy the Leibniz rule `D (a * b) = a * D b + D a * b`. diff --git a/src/ring_theory/integrally_closed.lean b/src/ring_theory/integrally_closed.lean index a2c2f43ab1a6c..874aaff7990da 100644 --- a/src/ring_theory/integrally_closed.lean +++ b/src/ring_theory/integrally_closed.lean @@ -9,6 +9,9 @@ import ring_theory.localization.integral /-! # Integrally closed rings +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + An integrally closed domain `R` contains all the elements of `Frac(R)` that are integral over `R`. A special case of integrally closed domains are the Dedekind domains. diff --git a/src/ring_theory/polynomial/rational_root.lean b/src/ring_theory/polynomial/rational_root.lean index 39cdfab524d69..67b786831fd8a 100644 --- a/src/ring_theory/polynomial/rational_root.lean +++ b/src/ring_theory/polynomial/rational_root.lean @@ -11,6 +11,9 @@ import ring_theory.polynomial.scale_roots /-! # Rational root theorem and integral root theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains the rational root theorem and integral root theorem. The rational root theorem for a unique factorization domain `A` with localization `S`, states that the roots of `p : A[X]` in `A`'s diff --git a/src/ring_theory/valuation/integral.lean b/src/ring_theory/valuation/integral.lean index 4969e3647a99a..147393bd09a4c 100644 --- a/src/ring_theory/valuation/integral.lean +++ b/src/ring_theory/valuation/integral.lean @@ -10,6 +10,9 @@ import ring_theory.valuation.integers /-! # Integral elements over the ring of integers of a valution +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The ring of integers is integrally closed inside the original ring. -/ diff --git a/src/topology/homotopy/H_spaces.lean b/src/topology/homotopy/H_spaces.lean index 722ad7fd54618..be4fc947d1914 100644 --- a/src/topology/homotopy/H_spaces.lean +++ b/src/topology/homotopy/H_spaces.lean @@ -9,6 +9,9 @@ import topology.homotopy.path /-! # H-spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines H-spaces mainly following the approach proposed by Serre in his paper *Homologie singulière des espaces fibrés*. The idea beaneath `H-spaces` is that they are topological spaces with a binary operation `⋀ : X → X → X` that is a homotopic-theoretic weakening of an From 5c1efce12ba86d4901463f61019832f6a4b1a0d0 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Sun, 4 Jun 2023 05:46:16 +0000 Subject: [PATCH 1139/1291] chore(*): add mathlib4 synchronization comments (#19151) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.jordan.basic` * `algebra.lie.abelian` * `algebra.lie.ideal_operations` * `algebra.lie.matrix` * `algebra.lie.of_associative` * `analysis.calculus.fderiv_symmetric` * `analysis.normed_space.dual_number` * `analysis.normed_space.triv_sq_zero_ext` * `category_theory.adhesive` * `data.mv_polynomial.pderiv` * `field_theory.chevalley_warning` * `linear_algebra.pi_tensor_product` * `number_theory.fermat_psp` * `number_theory.liouville.liouville_number` * `ring_theory.bezout` * `ring_theory.dedekind_domain.basic` * `ring_theory.derivation.lie` * `ring_theory.derivation.to_square_zero` * `topology.sheaves.sheafify` --- src/algebra/jordan/basic.lean | 3 +++ src/algebra/lie/abelian.lean | 3 +++ src/algebra/lie/ideal_operations.lean | 3 +++ src/algebra/lie/matrix.lean | 3 +++ src/algebra/lie/of_associative.lean | 3 +++ src/analysis/calculus/fderiv_symmetric.lean | 3 +++ src/analysis/normed_space/dual_number.lean | 3 +++ src/analysis/normed_space/triv_sq_zero_ext.lean | 3 +++ src/category_theory/adhesive.lean | 3 +++ src/data/mv_polynomial/pderiv.lean | 3 +++ src/field_theory/chevalley_warning.lean | 3 +++ src/linear_algebra/pi_tensor_product.lean | 3 +++ src/number_theory/fermat_psp.lean | 3 +++ src/number_theory/liouville/liouville_number.lean | 3 +++ src/ring_theory/bezout.lean | 3 +++ src/ring_theory/dedekind_domain/basic.lean | 3 +++ src/ring_theory/derivation/lie.lean | 3 +++ src/ring_theory/derivation/to_square_zero.lean | 3 +++ src/topology/sheaves/sheafify.lean | 3 +++ 19 files changed, 57 insertions(+) diff --git a/src/algebra/jordan/basic.lean b/src/algebra/jordan/basic.lean index a7259c6b0e11e..2c239fa99c6de 100644 --- a/src/algebra/jordan/basic.lean +++ b/src/algebra/jordan/basic.lean @@ -8,6 +8,9 @@ import algebra.lie.of_associative /-! # Jordan rings +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Let `A` be a non-unital, non-associative ring. Then `A` is said to be a (commutative, linear) Jordan ring if the multiplication is commutative and satisfies a weak associativity law known as the Jordan Identity: for all `a` and `b` in `A`, diff --git a/src/algebra/lie/abelian.lean b/src/algebra/lie/abelian.lean index 0b9c476b61cee..3464b04c602e3 100644 --- a/src/algebra/lie/abelian.lean +++ b/src/algebra/lie/abelian.lean @@ -9,6 +9,9 @@ import algebra.lie.ideal_operations /-! # Trivial Lie modules and Abelian Lie algebras +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The action of a Lie algebra `L` on a module `M` is trivial if `⁅x, m⁆ = 0` for all `x ∈ L` and `m ∈ M`. In the special case that `M = L` with the adjoint action, triviality corresponds to the concept of an Abelian Lie algebra. diff --git a/src/algebra/lie/ideal_operations.lean b/src/algebra/lie/ideal_operations.lean index a421f34acaa43..b1ab8c6d38e28 100644 --- a/src/algebra/lie/ideal_operations.lean +++ b/src/algebra/lie/ideal_operations.lean @@ -8,6 +8,9 @@ import algebra.lie.submodule /-! # Ideal operations for Lie algebras +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given a Lie module `M` over a Lie algebra `L`, there is a natural action of the Lie ideals of `L` on the Lie submodules of `M`. In the special case that `M = L` with the adjoint action, this provides a pairing of Lie ideals which is especially important. For example, it can be used to diff --git a/src/algebra/lie/matrix.lean b/src/algebra/lie/matrix.lean index eb02bbf72b6b9..afd6b6dc5c42f 100644 --- a/src/algebra/lie/matrix.lean +++ b/src/algebra/lie/matrix.lean @@ -10,6 +10,9 @@ import linear_algebra.matrix.to_linear_equiv /-! # Lie algebras of matrices +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + An important class of Lie algebras are those arising from the associative algebra structure on square matrices over a commutative ring. This file provides some very basic definitions whose primary value stems from their utility when constructing the classical Lie algebras using matrices. diff --git a/src/algebra/lie/of_associative.lean b/src/algebra/lie/of_associative.lean index 5049af2c5ecd2..e63c1a71f10e1 100644 --- a/src/algebra/lie/of_associative.lean +++ b/src/algebra/lie/of_associative.lean @@ -11,6 +11,9 @@ import algebra.algebra.subalgebra.basic /-! # Lie algebras of associative algebras +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the Lie algebra structure that arises on an associative algebra via the ring commutator. diff --git a/src/analysis/calculus/fderiv_symmetric.lean b/src/analysis/calculus/fderiv_symmetric.lean index 52a36b3928461..c631d08764501 100644 --- a/src/analysis/calculus/fderiv_symmetric.lean +++ b/src/analysis/calculus/fderiv_symmetric.lean @@ -8,6 +8,9 @@ import analysis.calculus.mean_value /-! # Symmetry of the second derivative +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We show that, over the reals, the second derivative is symmetric. The most precise result is `convex.second_derivative_within_at_symmetric`. It asserts that, diff --git a/src/analysis/normed_space/dual_number.lean b/src/analysis/normed_space/dual_number.lean index fbe680a7fcad0..d0ffa285df06f 100644 --- a/src/analysis/normed_space/dual_number.lean +++ b/src/analysis/normed_space/dual_number.lean @@ -9,6 +9,9 @@ import analysis.normed_space.triv_sq_zero_ext /-! # Results on `dual_number R` related to the norm +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + These are just restatements of similar statements about `triv_sq_zero_ext R M`. ## Main results diff --git a/src/analysis/normed_space/triv_sq_zero_ext.lean b/src/analysis/normed_space/triv_sq_zero_ext.lean index a21d6f4182856..4f10169ef1883 100644 --- a/src/analysis/normed_space/triv_sq_zero_ext.lean +++ b/src/analysis/normed_space/triv_sq_zero_ext.lean @@ -10,6 +10,9 @@ import topology.instances.triv_sq_zero_ext /-! # Results on `triv_sq_zero_ext R M` related to the norm +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + For now, this file contains results about `exp` for this type. ## Main results diff --git a/src/category_theory/adhesive.lean b/src/category_theory/adhesive.lean index e9762d9c6f4e4..e253aa0893c9f 100644 --- a/src/category_theory/adhesive.lean +++ b/src/category_theory/adhesive.lean @@ -10,6 +10,9 @@ import category_theory.limits.shapes.kernel_pair # Adhesive categories +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions - `category_theory.is_pushout.is_van_kampen`: A convenience formulation for a pushout being a van Kampen colimit. diff --git a/src/data/mv_polynomial/pderiv.lean b/src/data/mv_polynomial/pderiv.lean index 1e95dac506240..25f20db06a227 100644 --- a/src/data/mv_polynomial/pderiv.lean +++ b/src/data/mv_polynomial/pderiv.lean @@ -10,6 +10,9 @@ import data.mv_polynomial.derivation /-! # Partial derivatives of polynomials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the notion of the formal *partial derivative* of a polynomial, the derivative with respect to a single variable. This derivative is not connected to the notion of derivative from analysis. diff --git a/src/field_theory/chevalley_warning.lean b/src/field_theory/chevalley_warning.lean index 3e0de4a8b463b..8fa213e4bbe9f 100644 --- a/src/field_theory/chevalley_warning.lean +++ b/src/field_theory/chevalley_warning.lean @@ -9,6 +9,9 @@ import field_theory.finite.basic /-! # The Chevalley–Warning theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains a proof of the Chevalley–Warning theorem. Throughout most of this file, `K` denotes a finite field and `q` is notation for the cardinality of `K`. diff --git a/src/linear_algebra/pi_tensor_product.lean b/src/linear_algebra/pi_tensor_product.lean index 9e0ef6666844b..11fbb0d2578f3 100644 --- a/src/linear_algebra/pi_tensor_product.lean +++ b/src/linear_algebra/pi_tensor_product.lean @@ -10,6 +10,9 @@ import linear_algebra.multilinear.tensor_product /-! # Tensor product of an indexed family of modules over commutative semirings +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define the tensor product of an indexed family `s : ι → Type*` of modules over commutative semirings. We denote this space by `⨂[R] i, s i` and define it as `free_add_monoid (R × Π i, s i)` quotiented by the appropriate equivalence relation. The treatment follows very closely that of the diff --git a/src/number_theory/fermat_psp.lean b/src/number_theory/fermat_psp.lean index a1860ef1866ea..cde8495f8b725 100644 --- a/src/number_theory/fermat_psp.lean +++ b/src/number_theory/fermat_psp.lean @@ -10,6 +10,9 @@ import order.filter.cofinite /-! # Fermat Pseudoprimes +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define Fermat pseudoprimes: composite numbers that pass the Fermat primality test. A natural number `n` passes the Fermat primality test to base `b` (and is therefore deemed a "probable prime") if `n` divides `b ^ (n - 1) - 1`. `n` is a Fermat pseudoprime to base `b` if `n` diff --git a/src/number_theory/liouville/liouville_number.lean b/src/number_theory/liouville/liouville_number.lean index 78d446cf58326..5af7ad7379eae 100644 --- a/src/number_theory/liouville/liouville_number.lean +++ b/src/number_theory/liouville/liouville_number.lean @@ -8,6 +8,9 @@ import number_theory.liouville.basic # Liouville constants +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains a construction of a family of Liouville numbers, indexed by a natural number $m$. The most important property is that they are examples of transcendental real numbers. This fact is recorded in `transcendental_liouville_number`. diff --git a/src/ring_theory/bezout.lean b/src/ring_theory/bezout.lean index f2e588af810d1..07bbc8d8352c3 100644 --- a/src/ring_theory/bezout.lean +++ b/src/ring_theory/bezout.lean @@ -11,6 +11,9 @@ import algebra.gcd_monoid.integrally_closed # Bézout rings +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A Bézout ring (Bezout ring) is a ring whose finitely generated ideals are principal. Notible examples include principal ideal rings, valuation rings, and the ring of algebraic integers. diff --git a/src/ring_theory/dedekind_domain/basic.lean b/src/ring_theory/dedekind_domain/basic.lean index c4ad4e9238297..7a4715fa81baf 100644 --- a/src/ring_theory/dedekind_domain/basic.lean +++ b/src/ring_theory/dedekind_domain/basic.lean @@ -9,6 +9,9 @@ import ring_theory.polynomial.rational_root /-! # Dedekind domains +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the notion of a Dedekind domain (or Dedekind ring), as a Noetherian integrally closed commutative ring of Krull dimension at most one. diff --git a/src/ring_theory/derivation/lie.lean b/src/ring_theory/derivation/lie.lean index fea858509675e..458b0340a8936 100644 --- a/src/ring_theory/derivation/lie.lean +++ b/src/ring_theory/derivation/lie.lean @@ -9,6 +9,9 @@ import ring_theory.derivation.basic /-! # Results +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + - `derivation.lie_algebra`: The `R`-derivations from `A` to `A` form an lie algebra over `R`. -/ diff --git a/src/ring_theory/derivation/to_square_zero.lean b/src/ring_theory/derivation/to_square_zero.lean index 80791e2b871da..d9d0d6aae2484 100644 --- a/src/ring_theory/derivation/to_square_zero.lean +++ b/src/ring_theory/derivation/to_square_zero.lean @@ -9,6 +9,9 @@ import ring_theory.ideal.quotient_operations /-! # Results +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + - `derivation_to_square_zero_equiv_lift`: The `R`-derivations from `A` into a square-zero ideal `I` of `B` corresponds to the lifts `A →ₐ[R] B` of the map `A →ₐ[R] B ⧸ I`. diff --git a/src/topology/sheaves/sheafify.lean b/src/topology/sheaves/sheafify.lean index 3c707fbc5aa5e..7adc5d90beb87 100644 --- a/src/topology/sheaves/sheafify.lean +++ b/src/topology/sheaves/sheafify.lean @@ -9,6 +9,9 @@ import topology.sheaves.stalks /-! # Sheafification of `Type` valued presheaves +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We construct the sheafification of a `Type` valued presheaf, as the subsheaf of dependent functions into the stalks consisting of functions which are locally germs. From 5f25c089cb34db4db112556f23c50d12da81b297 Mon Sep 17 00:00:00 2001 From: damiano Date: Sun, 4 Jun 2023 11:07:50 +0000 Subject: [PATCH 1140/1291] chore(archive/imo): change namespace `imo` to `imoYYYY_qX` (#19152) This PR fixes a mistake that I made when namespacing files in `archive.imo`. Now, the files whose namespace was `imo` have namespace their file name (minus .lean). Changes: * namespace `imo` changed to `imoYYYY_qX`; * added `nolint dup_namespace` tags to the declarations that have a duplication; * changed namespace `imo_1987_q1` to `imo1987_q1` for consistency. [Zulip discussion](https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/mathport/near/363379951) Co-authored-by: Scott Morrison --- archive/imo/imo1959_q1.lean | 8 +++++--- archive/imo/imo1962_q4.lean | 13 ++++++++++--- archive/imo/imo1972_q5.lean | 4 ++-- archive/imo/imo1975_q1.lean | 4 ++-- archive/imo/imo1977_q6.lean | 7 ++++--- archive/imo/imo1987_q1.lean | 4 ++-- archive/imo/imo1988_q6.lean | 9 +++++---- archive/imo/imo1994_q1.lean | 7 ++++--- archive/imo/imo1998_q2.lean | 12 +++++++----- archive/imo/imo2001_q2.lean | 7 ++++--- archive/imo/imo2001_q6.lean | 4 ---- archive/imo/imo2005_q3.lean | 7 ++++--- archive/imo/imo2005_q4.lean | 4 ++-- archive/imo/imo2006_q3.lean | 9 +++++---- archive/imo/imo2006_q5.lean | 7 ++++--- archive/imo/imo2008_q2.lean | 4 ++-- archive/imo/imo2008_q3.lean | 7 ++++--- archive/imo/imo2008_q4.lean | 7 ++++--- archive/imo/imo2011_q3.lean | 4 ---- archive/imo/imo2011_q5.lean | 4 ---- archive/imo/imo2013_q1.lean | 7 ++++--- archive/imo/imo2013_q5.lean | 7 ++++--- archive/imo/imo2019_q1.lean | 4 ++-- archive/imo/imo2019_q2.lean | 13 +++++++------ archive/imo/imo2019_q4.lean | 7 ++++--- archive/imo/imo2020_q2.lean | 4 ---- archive/imo/imo2021_q1.lean | 4 ++-- 27 files changed, 93 insertions(+), 85 deletions(-) diff --git a/archive/imo/imo1959_q1.lean b/archive/imo/imo1959_q1.lean index b1af08fd9b9fa..bb0c5f8b63b36 100644 --- a/archive/imo/imo1959_q1.lean +++ b/archive/imo/imo1959_q1.lean @@ -18,7 +18,7 @@ as saying the numerator and denominator are relatively prime. open nat -namespace imo +namespace imo1959_q1 lemma calculation (n k : ℕ) (h1 : k ∣ 21 * n + 4) (h2 : k ∣ 14 * n + 3) : k ∣ 1 := have h3 : k ∣ 2 * (21 * n + 4), from h1.mul_left 2, @@ -26,7 +26,9 @@ have h4 : k ∣ 3 * (14 * n + 3), from h2.mul_left 3, have h5 : 3 * (14 * n + 3) = 2 * (21 * n + 4) + 1, by ring, (nat.dvd_add_right h3).mp (h5 ▸ h4) +end imo1959_q1 + +open imo1959_q1 + theorem imo1959_q1 : ∀ n : ℕ, coprime (21 * n + 4) (14 * n + 3) := assume n, coprime_of_dvd' $ λ k hp h1 h2, calculation n k h1 h2 - -end imo diff --git a/archive/imo/imo1962_q4.lean b/archive/imo/imo1962_q4.lean index 65b534aba58e9..95b725d68beca 100644 --- a/archive/imo/imo1962_q4.lean +++ b/archive/imo/imo1962_q4.lean @@ -17,7 +17,7 @@ in fact the simplest form of the set of solutions, and then prove it equals the open real open_locale real -namespace imo +namespace imo1962_q4 noncomputable theory @@ -82,6 +82,10 @@ begin split; intro; linarith end +end imo1962_q4 + +open imo1962_q4 + /- The final theorem is now just gluing together our lemmas. -/ @@ -92,6 +96,7 @@ begin exact exists_or_distrib.symm end +namespace imo1962_q4 /- We now present a second solution. The key to this solution is that, when the identity is @@ -122,6 +127,10 @@ begin split; intro; linarith end +end imo1962_q4 + +open imo1962_q4 + /- Again, the final theorem is now just gluing together our lemmas. -/ @@ -131,5 +140,3 @@ calc problem_equation x ↔ cos x ^ 2 + cos (2 * x) ^ 2 + cos (3 * x) ^ 2 = 1 : by refl ... ↔ cos (2 * x) = 0 ∨ cos (3 * x) = 0 : by simp [cos_two_mul, cos_three_mul, formula] ... ↔ x ∈ solution_set : by { rw [solve_cos2x_0, solve_cos3x_0, ← exists_or_distrib], refl } - -end imo diff --git a/archive/imo/imo1972_q5.lean b/archive/imo/imo1972_q5.lean index 27b0e32c8e3f0..c796e589190f5 100644 --- a/archive/imo/imo1972_q5.lean +++ b/archive/imo/imo1972_q5.lean @@ -15,7 +15,7 @@ Problem: `f` and `g` are real-valued functions defined on the real line. For all Prove that `|g(x)| ≤ 1` for all `x`. -/ -namespace imo +namespace imo1972_q5 /-- This proof begins by introducing the supremum of `f`, `k ≤ 1` as well as `k' = k / ‖g y‖`. We then @@ -123,4 +123,4 @@ begin linarith, end -end imo +end imo1972_q5 diff --git a/archive/imo/imo1975_q1.lean b/archive/imo/imo1975_q1.lean index 43a6e9409a65b..0035c3fe5d698 100644 --- a/archive/imo/imo1975_q1.lean +++ b/archive/imo/imo1975_q1.lean @@ -31,7 +31,7 @@ variables (hx : antitone_on x (finset.Icc 1 n)) variables (hy : antitone_on y (finset.Icc 1 n)) include hx hy hσ -namespace imo +namespace imo1975_q1 theorem IMO_1975_Q1 : ∑ i in finset.Icc 1 n, (x i - y i) ^ 2 ≤ ∑ i in finset.Icc 1 n, (x i - y (σ i)) ^ 2 := @@ -49,4 +49,4 @@ begin exact antitone_on.monovary_on hx hy end -end imo +end imo1975_q1 diff --git a/archive/imo/imo1977_q6.lean b/archive/imo/imo1977_q6.lean index 79342518718c7..002795950f143 100644 --- a/archive/imo/imo1977_q6.lean +++ b/archive/imo/imo1977_q6.lean @@ -15,7 +15,7 @@ We first prove the problem statement for `f : ℕ → ℕ` then we use it to prove the statement for positive naturals. -/ -namespace imo +namespace imo1977_q6 theorem imo1977_q6_nat (f : ℕ → ℕ) (h : ∀ n, f (f n) < f (n + 1)) : ∀ n, f n = n := @@ -35,6 +35,9 @@ begin exact nat.eq_of_le_of_lt_succ (hf _) (hf_mono.lt_iff_lt.mp (h _)) end +end imo1977_q6 +open imo1977_q6 + theorem imo1977_q6 (f : ℕ+ → ℕ+) (h : ∀ n, f (f n) < f (n + 1)) : ∀ n, f n = n := begin @@ -44,5 +47,3 @@ begin { simp }, { simpa using h _ } } end - -end imo diff --git a/archive/imo/imo1987_q1.lean b/archive/imo/imo1987_q1.lean index b4809852dd33b..b73054ddc9101 100644 --- a/archive/imo/imo1987_q1.lean +++ b/archive/imo/imo1987_q1.lean @@ -27,7 +27,7 @@ variables (α : Type*) [fintype α] [decidable_eq α] open_locale big_operators nat open equiv fintype function finset (range sum_const) set (Iic) -namespace imo_1987_q1 +namespace imo1987_q1 /-- The set of pairs `(x : α, σ : perm α)` such that `σ x = x` is equivalent to the set of pairs `(x : α, σ : perm {x}ᶜ)`. -/ @@ -89,4 +89,4 @@ theorem main {n : ℕ} (hn : 1 ≤ n) : ∑ k in range (n + 1), k * p (fin n) k = n! := by rw [main₀, nat.mul_factorial_pred (zero_lt_one.trans_le hn)] -end imo_1987_q1 +end imo1987_q1 diff --git a/archive/imo/imo1988_q6.lean b/archive/imo/imo1988_q6.lean index 8913a0c8dc30a..1f2888fbfa544 100644 --- a/archive/imo/imo1988_q6.lean +++ b/archive/imo/imo1988_q6.lean @@ -25,11 +25,11 @@ To illustrate the technique, we also prove a similar result. -- open_locale classical -namespace imo - local attribute [instance] classical.prop_decidable local attribute [simp] sq +namespace imo1988_q6 + /-- Constant descent Vieta jumping. This proof technique allows one to prove an arbitrary proposition `claim`, @@ -185,6 +185,9 @@ begin -- Hence p' = (c, m_x) lies on the upper branch, and we are done. end +end imo1988_q6 +open imo1988_q6 + /--Question 6 of IMO1988. If a and b are two natural numbers such that a*b+1 divides a^2 + b^2, show that their quotient is a perfect square.-/ lemma imo1988_q6 {a b : ℕ} (h : (a*b+1) ∣ a^2 + b^2) : @@ -298,5 +301,3 @@ begin obtain rfl|rfl := (nat.dvd_prime nat.prime_two).mp y_dvd; apply mul_left_cancel₀, exacts [one_ne_zero, h.symm, two_ne_zero, h.symm] } } end - -end imo diff --git a/archive/imo/imo1994_q1.lean b/archive/imo/imo1994_q1.lean index ffc8762b0e0d4..7e6d6dc9b43cf 100644 --- a/archive/imo/imo1994_q1.lean +++ b/archive/imo/imo1994_q1.lean @@ -32,7 +32,7 @@ open_locale big_operators open finset -namespace imo +namespace imo1994_q1 lemma tedious (m : ℕ) (k : fin (m+1)) : m - (m + (m + 1 - ↑k)) % (m + 1) = ↑k := begin @@ -45,6 +45,9 @@ begin linarith end +end imo1994_q1 +open imo1994_q1 + theorem imo1994_q1 (n : ℕ) (m : ℕ) (A : finset ℕ) (hm : A.card = m + 1) (hrange : ∀ a ∈ A, 0 < a ∧ a ≤ n) (hadd : ∀ (a b ∈ A), a + b ≤ n → a + b ∈ A) : (m + 1) * (n + 1) ≤ 2 * ∑ x in A, x := @@ -99,5 +102,3 @@ begin -- A set of size `k+1` embed in one of size `k`, which yields a contradiction simpa [fin.coe_sub, tedious] using card_le_of_subset hf, end - -end imo diff --git a/archive/imo/imo1998_q2.lean b/archive/imo/imo1998_q2.lean index 928d78975e0d5..25557c40c8595 100644 --- a/archive/imo/imo1998_q2.lean +++ b/archive/imo/imo1998_q2.lean @@ -38,7 +38,10 @@ Rearranging gives the result. -/ open_locale classical -namespace imo + +variables {C J : Type*} (r : C → J → Prop) + +namespace imo1998_q2 noncomputable theory @@ -48,8 +51,6 @@ abbreviation judge_pair (J : Type*) := J × J /-- A triple consisting of contestant together with an ordered pair of judges. -/ abbreviation agreed_triple (C J : Type*) := C × (judge_pair J) -variables {C J : Type*} (r : C → J → Prop) - /-- The first judge from an ordered pair of judges. -/ abbreviation judge_pair.judge₁ : judge_pair J → J := prod.fst @@ -196,6 +197,9 @@ lemma clear_denominators {a b k : ℕ} (ha : 0 < a) (hb : 0 < b) : (b - 1 : ℚ) / (2 * b) ≤ k / a ↔ (b - 1) * a ≤ k * (2 * b) := by rw div_le_div_iff; norm_cast; simp [ha, hb] +end imo1998_q2 +open imo1998_q2 + theorem imo1998_q2 [fintype J] [fintype C] (a b k : ℕ) (hC : fintype.card C = a) (hJ : fintype.card J = b) (ha : 0 < a) (hb : odd b) (hk : ∀ (p : judge_pair J), p.distinct → (agreed_contestants r p).card ≤ k) : @@ -214,5 +218,3 @@ begin { simp, }, { exact le_of_mul_le_mul_right h z.succ_pos, }, end - -end imo diff --git a/archive/imo/imo2001_q2.lean b/archive/imo/imo2001_q2.lean index bce2a35ab9b3e..51ca0c9aec19e 100644 --- a/archive/imo/imo2001_q2.lean +++ b/archive/imo/imo2001_q2.lean @@ -30,7 +30,7 @@ open real variables {a b c : ℝ} -namespace imo +namespace imo2001_q2 lemma denom_pos (ha : 0 < a) (hb : 0 < b) (hc : 0 < c) : 0 < a ^ 4 + b ^ 4 + c ^ 4 := @@ -66,6 +66,9 @@ have h₂ : c ^ 4 + a ^ 4 + b ^ 4 = a ^ 4 + b ^ 4 + c ^ 4, calc _ ≥ _ : add_le_add (add_le_add (bound ha hb hc) (bound hb hc ha)) (bound hc ha hb) ... = 1 : by rw [h₁, h₂, ← add_div, ← add_div, div_self $ ne_of_gt $ denom_pos ha hb hc] +end imo2001_q2 +open imo2001_q2 + theorem imo2001_q2 (ha : 0 < a) (hb : 0 < b) (hc : 0 < c) : 1 ≤ a / sqrt (a ^ 2 + 8 * b * c) + b / sqrt (b ^ 2 + 8 * c * a) + @@ -74,5 +77,3 @@ have h3 : ∀ {x : ℝ}, 0 < x → (x ^ (3 : ℝ)⁻¹) ^ 3 = x := λ x hx, show ↑3 = (3 : ℝ), by norm_num ▸ rpow_nat_inv_pow_nat hx.le three_ne_zero, calc 1 ≤ _ : imo2001_q2' (rpow_pos_of_pos ha _) (rpow_pos_of_pos hb _) (rpow_pos_of_pos hc _) ... = _ : by rw [h3 ha, h3 hb, h3 hc] - -end imo diff --git a/archive/imo/imo2001_q6.lean b/archive/imo/imo2001_q6.lean index 05f09cbd59265..c2fa9f829a855 100644 --- a/archive/imo/imo2001_q6.lean +++ b/archive/imo/imo2001_q6.lean @@ -19,8 +19,6 @@ Prove that $a*b + c*d$ is not prime. variables {a b c d : ℤ} -namespace imo - theorem imo2001_q6 (hd : 0 < d) (hdc : d < c) (hcb : c < b) (hba : b < a) (h : a*c + b*d = (a + b - c + d) * (-a + b + c + d)) : ¬ prime (a*b + c*d) := @@ -44,5 +42,3 @@ begin have : a*c + b*d ≤ a*d + b*c, { from int.le_of_dvd aux h2 }, nlinarith only [hba, hdc, h, this] }, end - -end imo diff --git a/archive/imo/imo2005_q3.lean b/archive/imo/imo2005_q3.lean index b963c26b3ac78..4cb5e03f1a9dd 100644 --- a/archive/imo/imo2005_q3.lean +++ b/archive/imo/imo2005_q3.lean @@ -18,7 +18,7 @@ factoring `(x^5-x^2)/(x^5+y^2+z^2) - (x^5-x^2)/(x^3*(x^2+y^2+z^2))` into a non-n and then making use of `xyz ≥ 1` to show `(x^5-x^2)/(x^3*(x^2+y^2+z^2)) ≥ (x^2-y*z)/(x^2+y^2+z^2)`. -/ -namespace imo +namespace imo2005_q3 lemma key_insight (x y z : ℝ) (hx : x > 0) (hy : y > 0) (hz : z > 0) (h : x*y*z ≥ 1) : (x^5-x^2)/(x^5+y^2+z^2) ≥ (x^2-y*z)/(x^2+y^2+z^2) := @@ -44,6 +44,9 @@ begin by { field_simp [h₂.ne', h₃.ne'], ring }, end +end imo2005_q3 +open imo2005_q3 + theorem imo2005_q3 (x y z : ℝ) (hx : x > 0) (hy : y > 0) (hz : z > 0) (h : x*y*z ≥ 1) : (x^5-x^2)/(x^5+y^2+z^2) + (y^5-y^2)/(y^5+z^2+x^2) + (z^5-z^2)/(z^5+x^2+y^2) ≥ 0 := begin @@ -61,5 +64,3 @@ begin (by linarith [sq_nonneg (x-y), sq_nonneg (y-z), sq_nonneg (z-x)]) (by linarith [sq_nonneg x, sq_nonneg y, sq_nonneg z]) }, end - -end imo diff --git a/archive/imo/imo2005_q4.lean b/archive/imo/imo2005_q4.lean index c31d8aeae822e..b403688ea1702 100644 --- a/archive/imo/imo2005_q4.lean +++ b/archive/imo/imo2005_q4.lean @@ -15,7 +15,7 @@ This is quite an easy problem, in which the key point is a modular arithmetic ca the sequence `a n` relative to an arbitrary prime. -/ -namespace imo +namespace imo2005_q4 /-- The sequence considered in the problem, `2 ^ n + 3 ^ n + 6 ^ n - 1`. -/ def a (n : ℕ) : ℤ := 2 ^ n + 3 ^ n + 6 ^ n - 1 @@ -96,4 +96,4 @@ begin contradiction, end -end imo +end imo2005_q4 diff --git a/archive/imo/imo2006_q3.lean b/archive/imo/imo2006_q3.lean index 9a266e1fb6a78..5d3378638a2e6 100644 --- a/archive/imo/imo2006_q3.lean +++ b/archive/imo/imo2006_q3.lean @@ -27,10 +27,10 @@ It involves making the substitution `x = a - b`, `y = b - c`, `z = c - a`, `s = a + b + c`. -/ -namespace imo - open real +namespace imo2006_q3 + /-- Replacing `x` and `y` with their average increases the left side. -/ lemma lhs_ineq {x y : ℝ} (hxy : 0 ≤ x * y) : 16 * x ^ 2 * y ^ 2 * (x + y) ^ 2 ≤ ((x + y) ^ 2) ^ 3 := @@ -135,11 +135,12 @@ begin { exact mul_nonneg (mul_nonneg (sq_nonneg _) zero_le_two) (sqrt_nonneg _) } end +end imo2006_q3 +open imo2006_q3 + theorem imo2006_q3 (M : ℝ) : (∀ a b c : ℝ, |a * b * (a^2 - b^2) + b * c * (b^2 - c^2) + c * a * (c^2 - a^2)| ≤ M * (a^2 + b^2 + c^2)^2) ↔ 9 * sqrt 2 / 32 ≤ M := ⟨proof₂ M, λ h _ _ _, le_trans proof₁ $ mul_le_mul_of_nonneg_right h $ sq_nonneg _⟩ - -end imo diff --git a/archive/imo/imo2006_q5.lean b/archive/imo/imo2006_q5.lean index e170e157a1798..1d0028ec3da52 100644 --- a/archive/imo/imo2006_q5.lean +++ b/archive/imo/imo2006_q5.lean @@ -41,7 +41,7 @@ $P(t)+t-a-b$, and we're again done. open function polynomial -namespace imo +namespace imo2006_q5 /-- If every entry in a cyclic list of integers divides the next, then they all have the same absolute value. -/ @@ -200,6 +200,9 @@ begin { rw ←ha, apply sub_dvd_eval_sub } } } end +end imo2006_q5 +open imo2006_q5 + /-- The general problem follows easily from the k = 2 case. -/ theorem imo2006_q5 {P : polynomial ℤ} (hP : 1 < P.nat_degree) {k : ℕ} (hk : 0 < k) : (P.comp^[k] X - X).roots.to_finset.card ≤ P.nat_degree := @@ -210,5 +213,3 @@ begin simp only [sub_eq_zero, is_root.def, eval_sub, iterate_comp_eval, eval_X] at ht, simpa [mem_roots hP', sub_eq_zero] using polynomial.is_periodic_pt_eval_two ⟨k, hk, ht⟩ end - -end imo diff --git a/archive/imo/imo2008_q2.lean b/archive/imo/imo2008_q2.lean index f633e57bbf8a9..eb4df1664207b 100644 --- a/archive/imo/imo2008_q2.lean +++ b/archive/imo/imo2008_q2.lean @@ -26,7 +26,7 @@ using `c`, `m` and `n`. We factor `LHS - 1` as a square, which finishes the proo set of rational solutions to the equation, and that `W` is infinite. -/ -namespace imo +namespace imo2008_q2 lemma subst_abc {x y z : ℝ} (h : x*y*z = 1) : ∃ a b c : ℝ, a ≠ 0 ∧ b ≠ 0 ∧ c ≠ 0 ∧ x = a/b ∧ y = b/c ∧ z = c /a := @@ -135,4 +135,4 @@ begin exact hW_inf.mono hW_sub_S, end -end imo +end imo2008_q2 diff --git a/archive/imo/imo2008_q3.lean b/archive/imo/imo2008_q3.lean index a86b5ab4fe95a..64e7080d53b99 100644 --- a/archive/imo/imo2008_q3.lean +++ b/archive/imo/imo2008_q3.lean @@ -28,7 +28,7 @@ Then `p = 2n + k ≥ 2n + √(p - 4) = 2n + √(2n + k - 4) > √(2n)` and we ar open real -namespace imo +namespace imo2008_q3 lemma p_lemma (p : ℕ) (hpp : nat.prime p) (hp_mod_4_eq_1 : p ≡ 1 [MOD 4]) (hp_gt_20 : p > 20) : ∃ n : ℕ, p ∣ n ^ 2 + 1 ∧ (p : ℝ) > 2 * n + sqrt(2 * n) := @@ -79,6 +79,9 @@ begin exact ⟨n, hnat₁, by linarith only [hreal₆, hreal₁]⟩, end +end imo2008_q3 +open imo2008_q3 + theorem imo2008_q3 : ∀ N : ℕ, ∃ n : ℕ, n ≥ N ∧ ∃ p : ℕ, nat.prime p ∧ p ∣ n ^ 2 + 1 ∧ (p : ℝ) > 2 * n + sqrt(2 * n) := begin @@ -92,5 +95,3 @@ begin exact ⟨n, hn_ge_N, p, hpp, hnat, hreal⟩, end - -end imo diff --git a/archive/imo/imo2008_q4.lean b/archive/imo/imo2008_q4.lean index cc9889a6dec1a..30d80008f57cb 100644 --- a/archive/imo/imo2008_q4.lean +++ b/archive/imo/imo2008_q4.lean @@ -23,12 +23,15 @@ The desired theorem is that either `f = λ x, x` or `f = λ x, 1/x` open real -namespace imo +namespace imo2008_q4 lemma abs_eq_one_of_pow_eq_one (x : ℝ) (n : ℕ) (hn : n ≠ 0) (h : x ^ n = 1) : |x| = 1 := by rw [← pow_left_inj (abs_nonneg x) zero_le_one (pos_iff_ne_zero.2 hn), one_pow, pow_abs, h, abs_one] +end imo2008_q4 +open imo2008_q4 + theorem imo2008_q4 (f : ℝ → ℝ) (H₁ : ∀ x > 0, f(x) > 0) : @@ -105,5 +108,3 @@ begin obtain ha₂ := abs_eq_one_of_pow_eq_one a 4 (show 4 ≠ 0, by norm_num) ha₁, rw abs_of_pos ha at ha₂, rw ha₂ at hfa₁, norm_num at hfa₁ }, end - -end imo diff --git a/archive/imo/imo2011_q3.lean b/archive/imo/imo2011_q3.lean index 1889e503ef54b..c3e5dfc5d8361 100644 --- a/archive/imo/imo2011_q3.lean +++ b/archive/imo/imo2011_q3.lean @@ -20,8 +20,6 @@ for all x and y. Prove that f(x) = 0 for all x ≤ 0. Direct translation of the solution found in https://www.imo-official.org/problems/IMO2011SL.pdf -/ -namespace imo - theorem imo2011_q3 (f : ℝ → ℝ) (hf : ∀ x y, f (x + y) ≤ y * f x + f (f x)) : ∀ x ≤ 0, f x = 0 := begin @@ -70,5 +68,3 @@ begin rw hno at hp, linarith }, end - -end imo diff --git a/archive/imo/imo2011_q5.lean b/archive/imo/imo2011_q5.lean index 5fe984f6114a6..abf6897ad352d 100644 --- a/archive/imo/imo2011_q5.lean +++ b/archive/imo/imo2011_q5.lean @@ -17,8 +17,6 @@ of positive integers. Suppose that, for any two integers open int -namespace imo - theorem imo2011_q5 (f : ℤ → ℤ) (hpos : ∀ n : ℤ, 0 < f n) (hdvd : ∀ m n : ℤ, f (m - n) ∣ f m - f n) : ∀ m n : ℤ, f m ≤ f n → f m ∣ f n := @@ -59,5 +57,3 @@ begin { -- m = n rw h_fm_eq_fn } end - -end imo diff --git a/archive/imo/imo2013_q1.lean b/archive/imo/imo2013_q1.lean index a2b0315d2f37e..90b7943ee58ed 100644 --- a/archive/imo/imo2013_q1.lean +++ b/archive/imo/imo2013_q1.lean @@ -27,7 +27,7 @@ We prove a slightly more general version where k does not need to be strictly po open_locale big_operators -namespace imo +namespace imo2013_q1 lemma arith_lemma (k n : ℕ) : 0 < 2 * n + 2^k.succ := calc 0 < 2 : zero_lt_two @@ -45,6 +45,9 @@ begin simp [finset.mem_range.mp hi] end +end imo2013_q1 +open imo2013_q1 + theorem imo2013_q1 (n : ℕ+) (k : ℕ) : (∃ m : ℕ → ℕ+, (1 : ℚ) + (2^k - 1) / n = (∏ i in finset.range k, (1 + 1 / m i))) := begin @@ -103,5 +106,3 @@ begin (1 + 1 / ↑(m pk)) : by rw [prod_lemma, hpm, ←hmpk, mul_comm] ... = ∏ i in finset.range pk.succ, (1 + 1 / m i) : by rw ← finset.prod_range_succ _ pk } end - -end imo diff --git a/archive/imo/imo2013_q5.lean b/archive/imo/imo2013_q5.lean index d27aa509953e8..7211de3e55827 100644 --- a/archive/imo/imo2013_q5.lean +++ b/archive/imo/imo2013_q5.lean @@ -29,7 +29,7 @@ https://www.imo-official.org/problems/IMO2013SL.pdf open_locale big_operators -namespace imo +namespace imo2013_q5 lemma le_of_all_pow_lt_succ {x y : ℝ} (hx : 1 < x) (hy : 1 < y) (h : ∀ n : ℕ, 0 < n → x^n - 1 < y^n) : @@ -198,6 +198,9 @@ begin linarith [H5 x hx, H5 _ h_big_enough] end +end imo2013_q5 +open imo2013_q5 + theorem imo2013_q5 (f : ℚ → ℝ) (H1 : ∀ x y, 0 < x → 0 < y → f (x * y) ≤ f x * f y) @@ -303,5 +306,3 @@ begin ... = (((x2num : ℚ) / (x2denom : ℚ) : ℚ) : ℝ) : by norm_cast ... = x : by rw ←hrat_expand2 end - -end imo diff --git a/archive/imo/imo2019_q1.lean b/archive/imo/imo2019_q1.lean index 46d3bffacab40..6fd6d266a6a21 100644 --- a/archive/imo/imo2019_q1.lean +++ b/archive/imo/imo2019_q1.lean @@ -19,7 +19,7 @@ Note that there is a much more compact proof of this fact in Isabelle/HOL - http://downthetypehole.de/paste/4YbGgqb4 -/ -namespace imo +namespace imo2019_q1 theorem imo2019Q1 (f : ℤ → ℤ) : (∀ a b : ℤ, f (2 * a) + 2 * (f b) = f (f (a + b))) ↔ @@ -52,4 +52,4 @@ begin { left, ext b, simpa [H, two_ne_zero] using H3 } end -end imo +end imo2019_q1 diff --git a/archive/imo/imo2019_q2.lean b/archive/imo/imo2019_q2.lean index 35d005480df5f..4289ec36cbd5a 100644 --- a/archive/imo/imo2019_q2.lean +++ b/archive/imo/imo2019_q2.lean @@ -56,10 +56,6 @@ rather than more literally with `affine_segment`. -/ library_note "IMO geometry formalization conventions" -namespace imo - -noncomputable theory - open affine affine.simplex euclidean_geometry finite_dimensional open_locale affine euclidean_geometry real @@ -70,6 +66,10 @@ variables [normed_add_comm_group V] [inner_product_space ℝ V] [metric_space Pt variables [normed_add_torsor V Pt] [hd2 : fact (finrank ℝ V = 2)] include hd2 +namespace imo2019_q2 + +noncomputable theory + /-- A configuration satisfying the conditions of the problem. We define this structure to avoid passing many hypotheses around as we build up information about the configuration; the final result for a statement of the problem not using this structure is then deduced from one in terms @@ -582,6 +582,9 @@ end end imo2019q2_cfg +end imo2019_q2 +open imo2019_q2 + theorem imo2019_q2 (A B C A₁ B₁ P Q P₁ Q₁ : Pt) (affine_independent_ABC : affine_independent ℝ ![A, B, C]) (wbtw_B_A₁_C : wbtw ℝ B A₁ C) (wbtw_A_B₁_C : wbtw ℝ A B₁ C) (wbtw_A_P_A₁ : wbtw ℝ A P A₁) @@ -593,5 +596,3 @@ theorem imo2019_q2 (A B C A₁ B₁ P Q P₁ Q₁ : Pt) (⟨A, B, C, A₁, B₁, P, Q, P₁, Q₁, affine_independent_ABC, wbtw_B_A₁_C, wbtw_A_B₁_C, wbtw_A_P_A₁, wbtw_B_Q_B₁, PQ_parallel_AB, P_ne_Q, sbtw_P_B₁_P₁, angle_PP₁C_eq_angle_BAC, C_ne_P₁, sbtw_Q_A₁_Q₁, angle_CQ₁Q_eq_angle_CBA, C_ne_Q₁⟩ : imo2019q2_cfg V Pt).result - -end imo diff --git a/archive/imo/imo2019_q4.lean b/archive/imo/imo2019_q4.lean index dfc00ea7b44d8..479df77d88099 100644 --- a/archive/imo/imo2019_q4.lean +++ b/archive/imo/imo2019_q4.lean @@ -26,7 +26,7 @@ individually. open_locale nat big_operators open finset multiplicity nat (hiding zero_le prime) -namespace imo +namespace imo2019_q4 theorem imo2019_q4_upper_bound {k n : ℕ} (hk : k > 0) (h : (k! : ℤ) = ∏ i in range n, (2 ^ n - 2 ^ i)) : n < 6 := @@ -69,6 +69,9 @@ begin convert add_le_add_left (add_le_add_left h5 (2 * n')) (n' * n') using 1, ring end +end imo2019_q4 +open imo2019_q4 + theorem imo2019_q4 {k n : ℕ} (hk : k > 0) (hn : n > 0) : (k! : ℤ) = (∏ i in range n, (2 ^ n - 2 ^ i)) ↔ (k, n) = (1, 1) ∨ (k, n) = (3, 2) := begin @@ -94,5 +97,3 @@ begin /- n = 5 -/ { refine monotone_factorial.ne_of_lt_of_lt_nat 10 _ _ _ h; norm_num }, end - -end imo diff --git a/archive/imo/imo2020_q2.lean b/archive/imo/imo2020_q2.lean index 4bc2f471ae33d..beea58a78ff17 100644 --- a/archive/imo/imo2020_q2.lean +++ b/archive/imo/imo2020_q2.lean @@ -21,8 +21,6 @@ the official solutions. open real -namespace imo - theorem imo2020_q2 (a b c d : ℝ) (hd0 : 0 < d) (hdc : d ≤ c) (hcb : c ≤ b) (hba : b ≤ a) (h1 : a + b + c + d = 1) : (a + 2 * b + 3 * c + 4 * d) * a ^ a * b ^ b * c ^ c * d ^ d < 1 := @@ -46,5 +44,3 @@ begin ... = (a + b + c + d) ^ 3 : by ring ... = 1 : by simp [h1] end - -end imo diff --git a/archive/imo/imo2021_q1.lean b/archive/imo/imo2021_q1.lean index 4a8f457c42722..04659d2cb76a2 100644 --- a/archive/imo/imo2021_q1.lean +++ b/archive/imo/imo2021_q1.lean @@ -43,7 +43,7 @@ which finishes the proof. open real -namespace imo +namespace imo2021_q1 lemma lower_bound (n l : ℕ) (hl : 2 + sqrt (4 + 2 * n) ≤ 2 * l) : n + 4 * l ≤ 2 * l * l := @@ -181,4 +181,4 @@ begin exact ⟨a, (hCA ha).2, b, (hCA hb).2, hab, h₁ a (hCA ha).1 b (hCA hb).1 hab⟩, end -end imo +end imo2021_q1 From 13361559d66b84f80b6d5a1c4a26aa5054766725 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Sun, 4 Jun 2023 16:20:36 +0000 Subject: [PATCH 1141/1291] chore(topology/sheaves/*): universe generalizations (#19153) Necessary but sadly insufficient for the request at https://leanprover.zulipchat.com/#narrow/stream/144837-PR-reviews/topic/.2319146.20sheaves.20on.20manifolds Co-authored-by: Scott Morrison --- .../morphisms/quasi_compact.lean | 4 +- .../morphisms/quasi_separated.lean | 6 +-- src/algebraic_geometry/ringed_space.lean | 4 +- src/algebraic_geometry/structure_sheaf.lean | 4 +- src/category_theory/limits/shapes/types.lean | 19 ++++++++- src/topology/sheaves/forget.lean | 28 +++++++------ src/topology/sheaves/local_predicate.lean | 40 +++++++++++-------- src/topology/sheaves/presheaf.lean | 20 +++++----- .../sheaves/presheaf_of_functions.lean | 8 ++-- .../sheaf_condition/unique_gluing.lean | 29 +++++++------- src/topology/sheaves/sheaf_of_functions.lean | 8 ++-- src/topology/sheaves/stalks.lean | 6 +-- 12 files changed, 102 insertions(+), 74 deletions(-) diff --git a/src/algebraic_geometry/morphisms/quasi_compact.lean b/src/algebraic_geometry/morphisms/quasi_compact.lean index b3281898a3059..5da4cb3432cfd 100644 --- a/src/algebraic_geometry/morphisms/quasi_compact.lean +++ b/src/algebraic_geometry/morphisms/quasi_compact.lean @@ -298,7 +298,7 @@ end /-- If `x : Γ(X, U)` is zero on `D(f)` for some `f : Γ(X, U)`, and `U` is quasi-compact, then `f ^ n * x = 0` for some `n`. -/ -lemma exists_pow_mul_eq_zero_of_res_basic_open_eq_zero_of_is_compact (X : Scheme) +lemma exists_pow_mul_eq_zero_of_res_basic_open_eq_zero_of_is_compact (X : Scheme.{u}) {U : opens X.carrier} (hU : is_compact U.1) (x f : X.presheaf.obj (op U)) (H : x |_ X.basic_open f = 0) : ∃ n : ℕ, f ^ n * x = 0 := @@ -322,7 +322,7 @@ begin use finset.univ.sup n, suffices : ∀ (i : s), X.presheaf.map (hom_of_le (h₁ i)).op (f ^ (finset.univ.sup n) * x) = 0, { subst e, - apply X.sheaf.eq_of_locally_eq (λ (i : s), (i : opens X.carrier)), + apply Top.sheaf.eq_of_locally_eq.{(u+1) u} X.sheaf (λ (i : s), (i : opens X.carrier)), intro i, rw map_zero, apply this }, diff --git a/src/algebraic_geometry/morphisms/quasi_separated.lean b/src/algebraic_geometry/morphisms/quasi_separated.lean index 8b75d385dcf82..f00888c66f0ba 100644 --- a/src/algebraic_geometry/morphisms/quasi_separated.lean +++ b/src/algebraic_geometry/morphisms/quasi_separated.lean @@ -345,7 +345,7 @@ begin exact e end -lemma exists_eq_pow_mul_of_is_compact_of_is_quasi_separated (X : Scheme) +lemma exists_eq_pow_mul_of_is_compact_of_is_quasi_separated (X : Scheme.{u}) (U : opens X.carrier) (hU : is_compact U.1) (hU' : is_quasi_separated U.1) (f : X.presheaf.obj (op U)) (x : X.presheaf.obj (op $ X.basic_open f)) : ∃ (n : ℕ) (y : X.presheaf.obj (op U)), y |_ X.basic_open f = (f |_ X.basic_open f) ^ n * x := @@ -401,7 +401,7 @@ begin ((X.presheaf.map (hom_of_le le_sup_left).op f) ^ (finset.univ.sup n + n₂) * y₁) = X.presheaf.map (hom_of_le $ inf_le_right).op ((X.presheaf.map (hom_of_le le_sup_right).op f) ^ (finset.univ.sup n + n₁) * y₂), - { fapply X.sheaf.eq_of_locally_eq' (λ i : s, i.1.1), + { fapply Top.sheaf.eq_of_locally_eq'.{(u+1) u} X.sheaf (λ i : s, i.1.1), { refine λ i, hom_of_le _, erw hs, exact le_supr _ _ }, { exact le_of_eq hs }, { intro i, @@ -419,7 +419,7 @@ begin -- By the sheaf condition, since `f ^ (n + n₂) * y₁ = f ^ (n + n₁) * y₂`, it can be glued into -- the desired section on `S ∪ U`. use (X.sheaf.obj_sup_iso_prod_eq_locus S U.1).inv ⟨⟨_ * _, _ * _⟩, this⟩, - refine X.sheaf.eq_of_locally_eq₂ + refine Top.sheaf.eq_of_locally_eq₂.{(u+1) u} X.sheaf (hom_of_le (_ : X.basic_open (X.presheaf.map (hom_of_le le_sup_left).op f) ≤ _)) (hom_of_le (_ : X.basic_open (X.presheaf.map (hom_of_le le_sup_right).op f) ≤ _)) _ _ _ _ _, { rw X.basic_open_res, exact inf_le_right }, diff --git a/src/algebraic_geometry/ringed_space.lean b/src/algebraic_geometry/ringed_space.lean index ff9354559ca56..268e8c7f83277 100644 --- a/src/algebraic_geometry/ringed_space.lean +++ b/src/algebraic_geometry/ringed_space.lean @@ -76,7 +76,7 @@ begin -- Let `g x` denote the inverse of `f` in `U x`. choose g hg using λ x : U, is_unit.exists_right_inv (h_unit x), -- We claim that these local inverses glue together to a global inverse of `f`. - obtain ⟨gl, gl_spec, -⟩ := X.sheaf.exists_unique_gluing' V U iVU hcover g _, + obtain ⟨gl, gl_spec, -⟩ := Top.sheaf.exists_unique_gluing'.{(v+1) v} X.sheaf V U iVU hcover g _, swap, { intros x y, apply section_ext X.sheaf (V x ⊓ V y), @@ -89,7 +89,7 @@ begin congr_arg (X.presheaf.germ (⟨z, hzVy⟩ : V y)) (hg y), ring_hom.map_one, ring_hom.map_one] }, apply is_unit_of_mul_eq_one f gl, - apply X.sheaf.eq_of_locally_eq' V U iVU hcover, + apply Top.sheaf.eq_of_locally_eq'.{(v+1) v} X.sheaf V U iVU hcover, intro i, rw [ring_hom.map_one, ring_hom.map_mul, gl_spec], exact hg i, diff --git a/src/algebraic_geometry/structure_sheaf.lean b/src/algebraic_geometry/structure_sheaf.lean index 1f2848001ed88..799927233cb81 100644 --- a/src/algebraic_geometry/structure_sheaf.lean +++ b/src/algebraic_geometry/structure_sheaf.lean @@ -772,10 +772,10 @@ begin rw to_basic_open_mk', -- Since the structure sheaf is a sheaf, we can show the desired equality locally. - -- Annoyingly, `sheaf.eq_of_locally_eq` requires an open cover indexed by a *type*, so we need to + -- Annoyingly, `sheaf.eq_of_locally_eq'` requires an open cover indexed by a *type*, so we need to -- coerce our finset `t` to a type first. let tt := ((t : set (basic_open f)) : Type u), - apply (structure_sheaf R).eq_of_locally_eq' + apply Top.sheaf.eq_of_locally_eq'.{(u+1) u} (structure_sheaf R) (λ i : tt, basic_open (h i)) (basic_open f) (λ i : tt, iDh i), { -- This feels a little redundant, since already have `ht_cover` as a hypothesis -- Unfortunately, `ht_cover` uses a bounded union over the set `t`, while here we have the diff --git a/src/category_theory/limits/shapes/types.lean b/src/category_theory/limits/shapes/types.lean index 4666f03ce1a24..6439eb4210fb4 100644 --- a/src/category_theory/limits/shapes/types.lean +++ b/src/category_theory/limits/shapes/types.lean @@ -51,13 +51,28 @@ local attribute [tidy] tactic.discrete_cases /-- A restatement of `types.lift_π_apply` that uses `pi.π` and `pi.lift`. -/ @[simp] lemma pi_lift_π_apply - {β : Type u} (f : β → Type u) {P : Type u} (s : Π b, P ⟶ f b) (b : β) (x : P) : + {β : Type v} (f : β → Type max v u) {P : Type max v u} (s : Π b, P ⟶ f b) (b : β) (x : P) : + (pi.π f b : (∏ f) → f b) (@pi.lift β _ _ f _ P s x) = s b x := +congr_fun (limit.lift_π (fan.mk P s) ⟨b⟩) x + +/-- A restatement of `types.lift_π_apply` that uses `pi.π` and `pi.lift`, +with specialized universes. -/ +@[simp] +lemma pi_lift_π_apply' + {β : Type v} (f : β → Type v) {P : Type v} (s : Π b, P ⟶ f b) (b : β) (x : P) : (pi.π f b : (∏ f) → f b) (@pi.lift β _ _ f _ P s x) = s b x := congr_fun (limit.lift_π (fan.mk P s) ⟨b⟩) x /-- A restatement of `types.map_π_apply` that uses `pi.π` and `pi.map`. -/ @[simp] -lemma pi_map_π_apply {β : Type u} {f g : β → Type u} (α : Π j, f j ⟶ g j) (b : β) (x) : +lemma pi_map_π_apply {β : Type v} {f g : β → Type max v u} (α : Π j, f j ⟶ g j) (b : β) (x) : + (pi.π g b : (∏ g) → g b) (pi.map α x) = α b ((pi.π f b : (∏ f) → f b) x) := +limit.map_π_apply _ _ _ + +/-- A restatement of `types.map_π_apply` that uses `pi.π` and `pi.map`, +with specialized universes. -/ +@[simp] +lemma pi_map_π_apply' {β : Type v} {f g : β → Type v} (α : Π j, f j ⟶ g j) (b : β) (x) : (pi.π g b : (∏ g) → g b) (pi.map α x) = α b ((pi.π f b : (∏ f) → f b) x) := limit.map_π_apply _ _ _ diff --git a/src/topology/sheaves/forget.lean b/src/topology/sheaves/forget.lean index 3fb52ea35b299..f6fbde37a2ce5 100644 --- a/src/topology/sheaves/forget.lean +++ b/src/topology/sheaves/forget.lean @@ -42,13 +42,13 @@ namespace sheaf_condition open sheaf_condition_equalizer_products -universes v u₁ u₂ +universes v u₁ u₂ w -variables {C : Type u₁} [category.{v} C] [has_limits C] -variables {D : Type u₂} [category.{v} D] [has_limits D] -variables (G : C ⥤ D) [preserves_limits G] -variables {X : Top.{v}} (F : presheaf C X) -variables {ι : Type v} (U : ι → opens X) +variables {C : Type u₁} [category.{v} C] [has_limits_of_size.{w w} C] +variables {D : Type u₂} [category.{v} D] [has_limits_of_size.{w w} D] +variables (G : C ⥤ D) [preserves_limits_of_size.{w w} G] +variables {X : Top.{w}} (F : presheaf C X) +variables {ι : Type w} (U : ι → opens X) local attribute [reducible] diagram left_res right_res @@ -57,7 +57,7 @@ When `G` preserves limits, the sheaf condition diagram for `F` composed with `G` naturally isomorphic to the sheaf condition diagram for `F ⋙ G`. -/ def diagram_comp_preserves_limits : - diagram F U ⋙ G ≅ diagram.{v} (F ⋙ G) U := + diagram F U ⋙ G ≅ diagram.{w} (F ⋙ G) U := begin fapply nat_iso.of_components, rintro ⟨j⟩, @@ -85,7 +85,7 @@ When `G` preserves limits, the image under `G` of the sheaf condition fork for ` is the sheaf condition fork for `F ⋙ G`, postcomposed with the inverse of the natural isomorphism `diagram_comp_preserves_limits`. -/ -def map_cone_fork : G.map_cone (fork.{v} F U) ≅ +def map_cone_fork : G.map_cone (fork.{w} F U) ≅ (cones.postcompose (diagram_comp_preserves_limits G F U).inv).obj (fork (F ⋙ G) U) := cones.ext (iso.refl _) (λ j, begin @@ -102,16 +102,17 @@ end) end sheaf_condition -universes v u₁ u₂ +universes v u₁ u₂ w open sheaf_condition sheaf_condition_equalizer_products variables {C : Type u₁} [category.{v} C] {D : Type u₂} [category.{v} D] variables (G : C ⥤ D) variables [reflects_isomorphisms G] -variables [has_limits C] [has_limits D] [preserves_limits G] +variables [has_limits_of_size.{w w} C] [has_limits_of_size.{w w} D] + [preserves_limits_of_size.{w w} G] -variables {X : Top.{v}} (F : presheaf C X) +variables {X : Top.{w}} (F : presheaf C X) /-- If `G : C ⥤ D` is a functor which reflects isomorphisms and preserves limits @@ -171,7 +172,7 @@ begin -- image under `G` of the equalizer cone for the sheaf condition diagram. let c := fork (F ⋙ G) U, obtain ⟨hc⟩ := S U, - let d := G.map_cone (equalizer.fork (left_res.{v} F U) (right_res F U)), + let d := G.map_cone (equalizer.fork (left_res.{w} F U) (right_res F U)), letI := preserves_smallest_limits_of_preserves_limits G, have hd : is_limit d := preserves_limit.preserves (limit.is_limit _), -- Since both of these are limit cones @@ -182,7 +183,8 @@ begin -- introduced above. let d' := (cones.postcompose (diagram_comp_preserves_limits G F U).hom).obj d, have hd' : is_limit d' := - (is_limit.postcompose_hom_equiv (diagram_comp_preserves_limits G F U : _) d).symm hd, + (is_limit.postcompose_hom_equiv + (diagram_comp_preserves_limits.{v u₁ u₂ w} G F U : _) d).symm hd, -- Now everything works: we verify that `f` really is a morphism between these cones: let f' : c ⟶ d' := fork.mk_hom (G.map f) diff --git a/src/topology/sheaves/local_predicate.lean b/src/topology/sheaves/local_predicate.lean index 04659ff17fd1b..3b682723589bf 100644 --- a/src/topology/sheaves/local_predicate.lean +++ b/src/topology/sheaves/local_predicate.lean @@ -38,12 +38,12 @@ to the types in the ambient type family. We give conditions sufficient to show that this map is injective and/or surjective. -/ -universe v +universes v w noncomputable theory variables {X : Top.{v}} -variables (T : X → Type v) +variables (T : X → Type w) open topological_space open opposite @@ -69,12 +69,12 @@ variables (X) Continuity is a "prelocal" predicate on functions to a fixed topological space `T`. -/ @[simps] -def continuous_prelocal (T : Top.{v}) : prelocal_predicate (λ x : X, T) := +def continuous_prelocal (T : Top.{w}) : prelocal_predicate (λ x : X, T) := { pred := λ U f, continuous f, res := λ U V i f h, continuous.comp h (opens.open_embedding_of_le i.le).continuous, } /-- Satisfying the inhabited linter. -/ -instance inhabited_prelocal_predicate (T : Top.{v}) : inhabited (prelocal_predicate (λ x : X, T)) := +instance inhabited_prelocal_predicate (T : Top.{w}) : inhabited (prelocal_predicate (λ x : X, T)) := ⟨continuous_prelocal X T⟩ variables {X} @@ -150,7 +150,7 @@ lemma prelocal_predicate.sheafify_of {T : X → Type v} {P : prelocal_predicate The subpresheaf of dependent functions on `X` satisfying the "pre-local" predicate `P`. -/ @[simps] -def subpresheaf_to_Types (P : prelocal_predicate T) : presheaf (Type v) X := +def subpresheaf_to_Types (P : prelocal_predicate T) : presheaf (Type (max v w)) X := { obj := λ U, { f : Π x : unop U, T x // P.pred f }, map := λ U V i f, ⟨λ x, f.1 (i.unop x), P.res i.unop f.1 f.2⟩ }. @@ -211,22 +211,27 @@ end subpresheaf_to_Types The subsheaf of the sheaf of all dependently typed functions satisfying the local predicate `P`. -/ @[simps] -def subsheaf_to_Types (P : local_predicate T) : sheaf (Type v) X := +def subsheaf_to_Types (P : local_predicate T) : sheaf (Type max v w) X := ⟨subpresheaf_to_Types P.to_prelocal_predicate, subpresheaf_to_Types.is_sheaf P⟩ +-- TODO There is further universe generalization to do here, +-- but it depends on more difficult universe generalization in `topology.sheaves.stalks`. +-- The aim is to replace `T'` with the more general `T` below. +variables {T' : X → Type v} + /-- There is a canonical map from the stalk to the original fiber, given by evaluating sections. -/ -def stalk_to_fiber (P : local_predicate T) (x : X) : - (subsheaf_to_Types P).presheaf.stalk x ⟶ T x := +def stalk_to_fiber (P : local_predicate T') (x : X) : + (subsheaf_to_Types P).presheaf.stalk x ⟶ T' x := begin refine colimit.desc _ - { X := T x, ι := { app := λ U f, _, naturality' := _ } }, + { X := T' x, ι := { app := λ U f, _, naturality' := _ } }, { exact f.1 ⟨x, (unop U).2⟩, }, { tidy, } end -@[simp] lemma stalk_to_fiber_germ (P : local_predicate T) (U : opens X) (x : U) (f) : +@[simp] lemma stalk_to_fiber_germ (P : local_predicate T') (U : opens X) (x : U) (f) : stalk_to_fiber P x ((subsheaf_to_Types P).presheaf.germ x f) = f.1 x := begin dsimp [presheaf.germ, stalk_to_fiber], @@ -239,8 +244,8 @@ end The `stalk_to_fiber` map is surjective at `x` if every point in the fiber `T x` has an allowed section passing through it. -/ -lemma stalk_to_fiber_surjective (P : local_predicate T) (x : X) - (w : ∀ (t : T x), ∃ (U : open_nhds x) (f : Π y : U.1, T y) (h : P.pred f), f ⟨x, U.2⟩ = t) : +lemma stalk_to_fiber_surjective (P : local_predicate T') (x : X) + (w : ∀ (t : T' x), ∃ (U : open_nhds x) (f : Π y : U.1, T' y) (h : P.pred f), f ⟨x, U.2⟩ = t) : function.surjective (stalk_to_fiber P x) := λ t, begin @@ -254,16 +259,16 @@ end The `stalk_to_fiber` map is injective at `x` if any two allowed sections which agree at `x` agree on some neighborhood of `x`. -/ -lemma stalk_to_fiber_injective (P : local_predicate T) (x : X) - (w : ∀ (U V : open_nhds x) (fU : Π y : U.1, T y) (hU : P.pred fU) - (fV : Π y : V.1, T y) (hV : P.pred fV) (e : fU ⟨x, U.2⟩ = fV ⟨x, V.2⟩), +lemma stalk_to_fiber_injective (P : local_predicate T') (x : X) + (w : ∀ (U V : open_nhds x) (fU : Π y : U.1, T' y) (hU : P.pred fU) + (fV : Π y : V.1, T' y) (hV : P.pred fV) (e : fU ⟨x, U.2⟩ = fV ⟨x, V.2⟩), ∃ (W : open_nhds x) (iU : W ⟶ U) (iV : W ⟶ V), ∀ (w : W.1), fU (iU w : U.1) = fV (iV w : V.1)) : function.injective (stalk_to_fiber P x) := λ tU tV h, begin -- We promise to provide all the ingredients of the proof later: let Q : - ∃ (W : (open_nhds x)ᵒᵖ) (s : Π w : (unop W).1, T w) (hW : P.pred s), + ∃ (W : (open_nhds x)ᵒᵖ) (s : Π w : (unop W).1, T' w) (hW : P.pred s), tU = (subsheaf_to_Types P).presheaf.germ ⟨x, (unop W).2⟩ ⟨s, hW⟩ ∧ tV = (subsheaf_to_Types P).presheaf.germ ⟨x, (unop W).2⟩ ⟨s, hW⟩ := _, { choose W s hW e using Q, @@ -283,6 +288,9 @@ begin colimit_sound iV.op (subtype.eq (funext w).symm)⟩, }, end +-- TODO to continue generalizing universes here, we will need to deal with the issue noted +-- at `presheaf_to_Top` (universe generalization for the Yoneda embedding). + /-- Some repackaging: the presheaf of functions satisfying `continuous_prelocal` is just the same thing as diff --git a/src/topology/sheaves/presheaf.lean b/src/topology/sheaves/presheaf.lean index e4c312be03265..ee4af1ab86ca9 100644 --- a/src/topology/sheaves/presheaf.lean +++ b/src/topology/sheaves/presheaf.lean @@ -221,7 +221,7 @@ def pushforward_map {X Y : Top.{w}} (f : X ⟶ Y) {ℱ 𝒢 : X.presheaf C} (α open category_theory.limits section pullback -variable [has_colimits C] +variable [has_colimits_of_size.{w w} C] noncomputable theory /-- @@ -231,17 +231,17 @@ This is defined in terms of left Kan extensions, which is just a fancy way of sa "take the colimits over the open sets whose preimage contains U". -/ @[simps] -def pullback_obj {X Y : Top.{v}} (f : X ⟶ Y) (ℱ : Y.presheaf C) : X.presheaf C := +def pullback_obj {X Y : Top.{w}} (f : X ⟶ Y) (ℱ : Y.presheaf C) : X.presheaf C := (Lan (opens.map f).op).obj ℱ /-- Pulling back along continuous maps is functorial. -/ -def pullback_map {X Y : Top.{v}} (f : X ⟶ Y) {ℱ 𝒢 : Y.presheaf C} (α : ℱ ⟶ 𝒢) : +def pullback_map {X Y : Top.{w}} (f : X ⟶ Y) {ℱ 𝒢 : Y.presheaf C} (α : ℱ ⟶ 𝒢) : pullback_obj f ℱ ⟶ pullback_obj f 𝒢 := (Lan (opens.map f).op).map α /-- If `f '' U` is open, then `f⁻¹ℱ U ≅ ℱ (f '' U)`. -/ @[simps] -def pullback_obj_obj_of_image_open {X Y : Top.{v}} (f : X ⟶ Y) (ℱ : Y.presheaf C) (U : opens X) +def pullback_obj_obj_of_image_open {X Y : Top.{w}} (f : X ⟶ Y) (ℱ : Y.presheaf C) (U : opens X) (H : is_open (f '' U)) : (pullback_obj f ℱ).obj (op U) ≅ ℱ.obj (op ⟨_, H⟩) := begin let x : costructured_arrow (opens.map f).op (op U) := begin @@ -263,7 +263,7 @@ begin end namespace pullback -variables {X Y : Top.{v}} (ℱ : Y.presheaf C) +variables {X Y : Top.{w}} (ℱ : Y.presheaf C) /-- The pullback along the identity is isomorphic to the original presheaf. -/ def id : pullback_obj (𝟙 _) ℱ ≅ ℱ := @@ -366,30 +366,30 @@ by simpa [pushforward_to_of_iso, equivalence.to_adjunction] end iso -variables (C) [has_colimits C] +variables (C) [has_colimits_of_size.{w w} C] /-- Pullback a presheaf on `Y` along a continuous map `f : X ⟶ Y`, obtaining a presheaf on `X`. -/ @[simps map_app] -def pullback {X Y : Top.{v}} (f : X ⟶ Y) : Y.presheaf C ⥤ X.presheaf C := Lan (opens.map f).op +def pullback {X Y : Top.{w}} (f : X ⟶ Y) : Y.presheaf C ⥤ X.presheaf C := Lan (opens.map f).op @[simp] lemma pullback_obj_eq_pullback_obj {C} [category C] [has_colimits C] {X Y : Top.{w}} (f : X ⟶ Y) (ℱ : Y.presheaf C) : (pullback C f).obj ℱ = pullback_obj f ℱ := rfl /-- The pullback and pushforward along a continuous map are adjoint to each other. -/ @[simps unit_app_app counit_app_app] -def pushforward_pullback_adjunction {X Y : Top.{v}} (f : X ⟶ Y) : +def pushforward_pullback_adjunction {X Y : Top.{w}} (f : X ⟶ Y) : pullback C f ⊣ pushforward C f := Lan.adjunction _ _ /-- Pulling back along a homeomorphism is the same as pushing forward along its inverse. -/ -def pullback_hom_iso_pushforward_inv {X Y : Top.{v}} (H : X ≅ Y) : +def pullback_hom_iso_pushforward_inv {X Y : Top.{w}} (H : X ≅ Y) : pullback C H.hom ≅ pushforward C H.inv := adjunction.left_adjoint_uniq (pushforward_pullback_adjunction C H.hom) (presheaf_equiv_of_iso C H.symm).to_adjunction /-- Pulling back along the inverse of a homeomorphism is the same as pushing forward along it. -/ -def pullback_inv_iso_pushforward_hom {X Y : Top.{v}} (H : X ≅ Y) : +def pullback_inv_iso_pushforward_hom {X Y : Top.{w}} (H : X ≅ Y) : pullback C H.inv ≅ pushforward C H.hom := adjunction.left_adjoint_uniq (pushforward_pullback_adjunction C H.inv) diff --git a/src/topology/sheaves/presheaf_of_functions.lean b/src/topology/sheaves/presheaf_of_functions.lean index 07212f6643439..a6c432a1dea51 100644 --- a/src/topology/sheaves/presheaf_of_functions.lean +++ b/src/topology/sheaves/presheaf_of_functions.lean @@ -28,7 +28,7 @@ We construct some simple examples of presheaves of functions on a topological sp is the presheaf of rings of continuous complex-valued functions on `X`. -/ -universes v u +universes v u w open category_theory open topological_space @@ -42,7 +42,7 @@ variables (X : Top.{v}) The presheaf of dependently typed functions on `X`, with fibres given by a type family `T`. There is no requirement that the functions are continuous, here. -/ -def presheaf_to_Types (T : X → Type v) : X.presheaf (Type v) := +def presheaf_to_Types (T : X → Type w) : X.presheaf (Type (max v w)) := { obj := λ U, Π x : (unop U), T x, map := λ U V i g, λ (x : unop V), g (i.unop x), map_id' := λ U, by { ext g ⟨x, hx⟩, refl }, @@ -68,7 +68,7 @@ There is no requirement that the functions are continuous, here. -- We don't use `@[simps]` to generate the projection lemmas here, -- as it turns out to be useful to have `presheaf_to_Type_map` -- written as an equality of functions (rather than being applied to some argument). -def presheaf_to_Type (T : Type v) : X.presheaf (Type v) := +def presheaf_to_Type (T : Type w) : X.presheaf (Type max v w) := { obj := λ U, (unop U) → T, map := λ U V i g, g ∘ i.unop, map_id' := λ U, by { ext g ⟨x, hx⟩, refl }, @@ -86,6 +86,8 @@ rfl /-- The presheaf of continuous functions on `X` with values in fixed target topological space `T`. -/ +-- TODO it may prove useful to generalize the universes here, +-- but the definition would need to change. def presheaf_to_Top (T : Top.{v}) : X.presheaf (Type v) := (opens.to_Top X).op ⋙ (yoneda.obj T) diff --git a/src/topology/sheaves/sheaf_condition/unique_gluing.lean b/src/topology/sheaves/sheaf_condition/unique_gluing.lean index 2cbe750fc2433..543920fcc9b69 100644 --- a/src/topology/sheaves/sheaf_condition/unique_gluing.lean +++ b/src/topology/sheaves/sheaf_condition/unique_gluing.lean @@ -48,9 +48,9 @@ open topological_space open topological_space.opens open opposite -universes u v +universes u v w -variables {C : Type u} [category.{v} C] [concrete_category.{v} C] +variables {C : Type u} [category.{max w v} C] [concrete_category.{max w v} C] namespace Top @@ -60,7 +60,7 @@ section local attribute [instance] concrete_category.has_coe_to_sort concrete_category.has_coe_to_fun -variables {X : Top.{v}} (F : presheaf C X) {ι : Type v} (U : ι → opens X) +variables {X : Top.{w}} (F : presheaf C X) {ι : Type w} (U : ι → opens X) /-- A family of sections `sf` is compatible, if the restrictions of `sf i` and `sf j` to `U i ⊓ U j` @@ -85,14 +85,14 @@ We prove this to be equivalent to the usual one below in `is_sheaf_iff_is_sheaf_unique_gluing` -/ def is_sheaf_unique_gluing : Prop := -∀ ⦃ι : Type v⦄ (U : ι → opens X) (sf : Π i : ι, F.obj (op (U i))), +∀ ⦃ι : Type w⦄ (U : ι → opens X) (sf : Π i : ι, F.obj (op (U i))), is_compatible F U sf → ∃! s : F.obj (op (supr U)), is_gluing F U sf s end section type_valued -variables {X : Top.{v}} (F : presheaf (Type v) X) {ι : Type v} (U : ι → opens X) +variables {X : Top.{w}} (F : presheaf (Type max w v) X) {ι : Type w} (U : ι → opens X) /-- For presheaves of types, terms of `pi_opens F U` are just families of sections. @@ -100,7 +100,7 @@ For presheaves of types, terms of `pi_opens F U` are just families of sections. def pi_opens_iso_sections_family : pi_opens F U ≅ Π i : ι, F.obj (op (U i)) := limits.is_limit.cone_point_unique_up_to_iso (limit.is_limit (discrete.functor (λ i : ι, F.obj (op (U i))))) - ((types.product_limit_cone.{v v} (λ i : ι, F.obj (op (U i)))).is_limit) + ((types.product_limit_cone.{w (max w v)} (λ i : ι, F.obj (op (U i)))).is_limit) /-- Under the isomorphism `pi_opens_iso_sections_family`, compatibility of sections is the same @@ -112,8 +112,8 @@ lemma compatible_iff_left_res_eq_right_res (sf : pi_opens F U) : begin split ; intros h, { ext ⟨i, j⟩, - rw [left_res, types.limit.lift_π_apply', fan.mk_π_app, - right_res, types.limit.lift_π_apply', fan.mk_π_app], + rw [left_res, types.limit.lift_π_apply, fan.mk_π_app, + right_res, types.limit.lift_π_apply, fan.mk_π_app], exact h i j, }, { intros i j, convert congr_arg (limits.pi.π (λ p : ι × ι, F.obj (op (U p.1 ⊓ U p.2))) (i,j)) h, @@ -132,7 +132,7 @@ lemma is_gluing_iff_eq_res (sf : pi_opens F U) (s : F.obj (op (supr U))): begin split ; intros h, { ext ⟨i⟩, - rw [res, types.limit.lift_π_apply', fan.mk_π_app], + rw [res, types.limit.lift_π_apply, fan.mk_π_app], exact h i, }, { intro i, convert congr_arg (limits.pi.π (λ i : ι, F.obj (op (U i))) i) h, @@ -209,8 +209,9 @@ section local attribute [instance] concrete_category.has_coe_to_sort concrete_category.has_coe_to_fun -variables [has_limits C] [reflects_isomorphisms (forget C)] [preserves_limits (forget C)] -variables {X : Top.{v}} (F : presheaf C X) {ι : Type v} (U : ι → opens X) +variables [has_limits_of_size.{w w} C] [reflects_isomorphisms (forget C)] + [preserves_limits_of_size.{w w} (forget C)] +variables {X : Top.{w}} (F : presheaf C X) {ι : Type w} (U : ι → opens X) /-- For presheaves valued in a concrete category, whose forgetful functor reflects isomorphisms and @@ -235,10 +236,10 @@ section local attribute [instance] concrete_category.has_coe_to_sort concrete_category.has_coe_to_fun -variables [has_limits C] [reflects_isomorphisms (concrete_category.forget C)] -variables [preserves_limits (concrete_category.forget C)] +variables [has_limits_of_size.{w w} C] [reflects_isomorphisms (concrete_category.forget C)] +variables [preserves_limits_of_size.{w w} (concrete_category.forget C)] -variables {X : Top.{v}} (F : sheaf C X) {ι : Type v} (U : ι → opens X) +variables {X : Top.{w}} (F : sheaf C X) {ι : Type w} (U : ι → opens X) /-- A more convenient way of obtaining a unique gluing of sections for a sheaf. diff --git a/src/topology/sheaves/sheaf_of_functions.lean b/src/topology/sheaves/sheaf_of_functions.lean index 8b3e8bcc0e660..1ffac5ae6b1d9 100644 --- a/src/topology/sheaves/sheaf_of_functions.lean +++ b/src/topology/sheaves/sheaf_of_functions.lean @@ -33,11 +33,11 @@ open category_theory.limits open topological_space open topological_space.opens -universe u +universes v u noncomputable theory -variables (X : Top.{u}) +variables (X : Top.{v}) open Top @@ -101,13 +101,13 @@ namespace Top The sheaf of not-necessarily-continuous functions on `X` with values in type family `T : X → Type u`. -/ -def sheaf_to_Types (T : X → Type u) : sheaf (Type u) X := +def sheaf_to_Types (T : X → Type u) : sheaf (Type max v u) X := ⟨presheaf_to_Types X T, presheaf.to_Types_is_sheaf _ _⟩ /-- The sheaf of not-necessarily-continuous functions on `X` with values in a type `T`. -/ -def sheaf_to_Type (T : Type u) : sheaf (Type u) X := +def sheaf_to_Type (T : Type u) : sheaf (Type max v u) X := ⟨presheaf_to_Type X T, presheaf.to_Type_is_sheaf _ _⟩ end Top diff --git a/src/topology/sheaves/stalks.lean b/src/topology/sheaves/stalks.lean index ea3a7426b187a..c8564219b13bb 100644 --- a/src/topology/sheaves/stalks.lean +++ b/src/topology/sheaves/stalks.lean @@ -419,7 +419,7 @@ begin choose V m i₁ i₂ heq using λ x : U, F.presheaf.germ_eq x.1 x.2 x.2 s t (h x), -- Since `F` is a sheaf, we can prove the equality locally, if we can show that these -- neighborhoods form a cover of `U`. - apply F.eq_of_locally_eq' V U i₁, + apply Top.sheaf.eq_of_locally_eq'.{u v v} F V U i₁, { intros x hxU, rw [opens.mem_supr], exact ⟨⟨x, hxU⟩, m ⟨x, hxU⟩⟩ }, @@ -488,9 +488,9 @@ begin rw [opens.mem_supr], exact ⟨⟨x, hxU⟩, mV ⟨x, hxU⟩⟩ }, -- Since `F` is a sheaf, we can glue all the local preimages together to get a global preimage. - obtain ⟨s, s_spec, -⟩ := F.exists_unique_gluing' V U iVU V_cover sf _, + obtain ⟨s, s_spec, -⟩ := Top.sheaf.exists_unique_gluing'.{u v v} F V U iVU V_cover sf _, { use s, - apply G.eq_of_locally_eq' V U iVU V_cover, + apply Top.sheaf.eq_of_locally_eq'.{u v v} G V U iVU V_cover, intro x, rw [← comp_apply, ← f.1.naturality, comp_apply, s_spec, heq] }, { intros x y, From 533f62f4dd62a5aad24a04326e6e787c8f7e98b1 Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Mon, 5 Jun 2023 01:10:26 +0000 Subject: [PATCH 1142/1291] chore(algebraic_geometry/open_immersion): split (#19149) Split `algebraic_geometry/open_immersion` into a 1000-line file which doesn't mention schemes and a 1000-line file which does. This should open more porting targets (the first file should be available for porting immediately), but also helps towards a reorganization I have in mind for after-port: moving the [pre]sheafed-space and [locally-]ringed-space material from `algebraic_geometry` into a new folder in `topology` or `geometry`. --- src/algebraic_geometry/AffineScheme.lean | 2 +- src/algebraic_geometry/gluing.lean | 1 + .../locally_ringed_space/has_colimits.lean | 2 +- .../Scheme.lean} | 1063 +--------------- .../open_immersion/basic.lean | 1081 +++++++++++++++++ .../presheafed_space/gluing.lean | 2 +- 6 files changed, 1089 insertions(+), 1062 deletions(-) rename src/algebraic_geometry/{open_immersion.lean => open_immersion/Scheme.lean} (50%) create mode 100644 src/algebraic_geometry/open_immersion/basic.lean diff --git a/src/algebraic_geometry/AffineScheme.lean b/src/algebraic_geometry/AffineScheme.lean index 916bf40550d9d..59201acbd1616 100644 --- a/src/algebraic_geometry/AffineScheme.lean +++ b/src/algebraic_geometry/AffineScheme.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Andrew Yang -/ import algebraic_geometry.Gamma_Spec_adjunction -import algebraic_geometry.open_immersion +import algebraic_geometry.open_immersion.Scheme import category_theory.limits.opposites import ring_theory.localization.inv_submonoid diff --git a/src/algebraic_geometry/gluing.lean b/src/algebraic_geometry/gluing.lean index 18c770363b79a..2a5345bbf6a65 100644 --- a/src/algebraic_geometry/gluing.lean +++ b/src/algebraic_geometry/gluing.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Andrew Yang -/ import algebraic_geometry.presheafed_space.gluing +import algebraic_geometry.open_immersion.Scheme /-! # Gluing Schemes diff --git a/src/algebraic_geometry/locally_ringed_space/has_colimits.lean b/src/algebraic_geometry/locally_ringed_space/has_colimits.lean index 34efd4dff6154..51c20a36cb793 100644 --- a/src/algebraic_geometry/locally_ringed_space/has_colimits.lean +++ b/src/algebraic_geometry/locally_ringed_space/has_colimits.lean @@ -5,7 +5,7 @@ Authors: Andrew Yang -/ import algebraic_geometry.locally_ringed_space import algebra.category.Ring.constructions -import algebraic_geometry.open_immersion +import algebraic_geometry.open_immersion.basic import category_theory.limits.constructions.limits_of_products_and_equalizers /-! diff --git a/src/algebraic_geometry/open_immersion.lean b/src/algebraic_geometry/open_immersion/Scheme.lean similarity index 50% rename from src/algebraic_geometry/open_immersion.lean rename to src/algebraic_geometry/open_immersion/Scheme.lean index 87d243476e371..571f52b4d17c3 100644 --- a/src/algebraic_geometry/open_immersion.lean +++ b/src/algebraic_geometry/open_immersion/Scheme.lean @@ -3,61 +3,17 @@ Copyright (c) 2021 Andrew Yang. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Andrew Yang -/ -import algebraic_geometry.presheafed_space.has_colimits -import category_theory.limits.shapes.binary_products -import category_theory.limits.preserves.shapes.pullbacks -import topology.sheaves.functors -import topology.category.Top.limits.pullbacks +import algebraic_geometry.open_immersion.basic import algebraic_geometry.Scheme -import category_theory.limits.shapes.strict_initial import category_theory.limits.shapes.comm_sq -import algebra.category.Ring.instances /-! -# Open immersions of structured spaces - -We say that a morphism of presheafed spaces `f : X ⟶ Y` is an open immersions if -the underlying map of spaces is an open embedding `f : X ⟶ U ⊆ Y`, -and the sheaf map `Y(V) ⟶ f _* X(V)` is an iso for each `V ⊆ U`. - -Abbreviations are also provided for `SheafedSpace`, `LocallyRingedSpace` and `Scheme`. - -## Main definitions - -* `algebraic_geometry.PresheafedSpace.is_open_immersion`: the `Prop`-valued typeclass asserting - that a PresheafedSpace hom `f` is an open_immersion. -* `algebraic_geometry.is_open_immersion`: the `Prop`-valued typeclass asserting - that a Scheme morphism `f` is an open_immersion. -* `algebraic_geometry.PresheafedSpace.is_open_immersion.iso_restrict`: The source of an - open immersion is isomorphic to the restriction of the target onto the image. -* `algebraic_geometry.PresheafedSpace.is_open_immersion.lift`: Any morphism whose range is - contained in an open immersion factors though the open immersion. -* `algebraic_geometry.PresheafedSpace.is_open_immersion.to_SheafedSpace`: If `f : X ⟶ Y` is an - open immersion of presheafed spaces, and `Y` is a sheafed space, then `X` is also a sheafed - space. The morphism as morphisms of sheafed spaces is given by `to_SheafedSpace_hom`. -* `algebraic_geometry.PresheafedSpace.is_open_immersion.to_LocallyRingedSpace`: If `f : X ⟶ Y` is - an open immersion of presheafed spaces, and `Y` is a locally ringed space, then `X` is also a - locally ringed space. The morphism as morphisms of locally ringed spaces is given by - `to_LocallyRingedSpace_hom`. - -## Main results - -* `algebraic_geometry.PresheafedSpace.is_open_immersion.comp`: The composition of two open - immersions is an open immersion. -* `algebraic_geometry.PresheafedSpace.is_open_immersion.of_iso`: An iso is an open immersion. -* `algebraic_geometry.PresheafedSpace.is_open_immersion.to_iso`: - A surjective open immersion is an isomorphism. -* `algebraic_geometry.PresheafedSpace.is_open_immersion.stalk_iso`: An open immersion induces - an isomorphism on stalks. -* `algebraic_geometry.PresheafedSpace.is_open_immersion.has_pullback_of_left`: If `f` is an open - immersion, then the pullback `(f, g)` exists (and the forgetful functor to `Top` preserves it). -* `algebraic_geometry.PresheafedSpace.is_open_immersion.pullback_snd_of_left`: Open immersions - are stable under pullbacks. -* `algebraic_geometry.SheafedSpace.is_open_immersion.of_stalk_iso` An (topological) open embedding - between two sheafed spaces is an open immersion if all the stalk maps are isomorphisms. +# Open immersions of schemes -/ +noncomputable theory + open topological_space category_theory opposite open category_theory.limits namespace algebraic_geometry @@ -66,28 +22,6 @@ universes v v₁ v₂ u variables {C : Type u} [category.{v} C] -/-- -An open immersion of PresheafedSpaces is an open embedding `f : X ⟶ U ⊆ Y` of the underlying -spaces, such that the sheaf map `Y(V) ⟶ f _* X(V)` is an iso for each `V ⊆ U`. --/ -class PresheafedSpace.is_open_immersion {X Y : PresheafedSpace.{v} C} (f : X ⟶ Y) : Prop := -(base_open : open_embedding f.base) -(c_iso : ∀ U : opens X, is_iso (f.c.app (op (base_open.is_open_map.functor.obj U)))) - -/-- -A morphism of SheafedSpaces is an open immersion if it is an open immersion as a morphism -of PresheafedSpaces --/ -abbreviation SheafedSpace.is_open_immersion {X Y : SheafedSpace.{v} C} (f : X ⟶ Y) : Prop := -PresheafedSpace.is_open_immersion f - -/-- -A morphism of LocallyRingedSpaces is an open immersion if it is an open immersion as a morphism -of SheafedSpaces --/ -abbreviation LocallyRingedSpace.is_open_immersion {X Y : LocallyRingedSpace} (f : X ⟶ Y) : Prop := -SheafedSpace.is_open_immersion f.1 - /-- A morphism of Schemes is an open immersion if it is an open immersion as a morphism of LocallyRingedSpaces @@ -95,997 +29,8 @@ of LocallyRingedSpaces abbreviation is_open_immersion {X Y : Scheme} (f : X ⟶ Y) : Prop := LocallyRingedSpace.is_open_immersion f -namespace PresheafedSpace.is_open_immersion - -open PresheafedSpace - -local notation `is_open_immersion` := PresheafedSpace.is_open_immersion - -attribute [instance] is_open_immersion.c_iso - -section - -variables {X Y : PresheafedSpace.{v} C} {f : X ⟶ Y} (H : is_open_immersion f) - -/-- The functor `opens X ⥤ opens Y` associated with an open immersion `f : X ⟶ Y`. -/ -abbreviation open_functor := H.base_open.is_open_map.functor - -/-- An open immersion `f : X ⟶ Y` induces an isomorphism `X ≅ Y|_{f(X)}`. -/ -@[simps hom_c_app] noncomputable -def iso_restrict : X ≅ Y.restrict H.base_open := -PresheafedSpace.iso_of_components (iso.refl _) -begin - symmetry, - fapply nat_iso.of_components, - intro U, - refine as_iso (f.c.app (op (H.open_functor.obj (unop U)))) ≪≫ X.presheaf.map_iso (eq_to_iso _), - { induction U using opposite.rec, - cases U, - dsimp only [is_open_map.functor, functor.op, opens.map], - congr' 2, - erw set.preimage_image_eq _ H.base_open.inj, - refl }, - { intros U V i, - simp only [category_theory.eq_to_iso.hom, Top.presheaf.pushforward_obj_map, category.assoc, - functor.op_map, iso.trans_hom, as_iso_hom, functor.map_iso_hom, ←X.presheaf.map_comp], - erw [f.c.naturality_assoc, ←X.presheaf.map_comp], - congr } -end - -@[simp] lemma iso_restrict_hom_of_restrict : H.iso_restrict.hom ≫ Y.of_restrict _ = f := -begin - ext, - { simp only [comp_c_app, iso_restrict_hom_c_app, nat_trans.comp_app, - eq_to_hom_refl, of_restrict_c_app, category.assoc, whisker_right_id'], - erw [category.comp_id, f.c.naturality_assoc, ←X.presheaf.map_comp], - transitivity f.c.app x ≫ X.presheaf.map (𝟙 _), - { congr }, - { erw [X.presheaf.map_id, category.comp_id] } }, - { refl, } -end - -@[simp] lemma iso_restrict_inv_of_restrict : H.iso_restrict.inv ≫ f = Y.of_restrict _ := -by { rw [iso.inv_comp_eq, iso_restrict_hom_of_restrict] } - -instance mono [H : is_open_immersion f] : mono f := -by { rw ← H.iso_restrict_hom_of_restrict, apply mono_comp } - -/-- The composition of two open immersions is an open immersion. -/ -instance comp {Z : PresheafedSpace C} (f : X ⟶ Y) [hf : is_open_immersion f] (g : Y ⟶ Z) - [hg : is_open_immersion g] : - is_open_immersion (f ≫ g) := -{ base_open := hg.base_open.comp hf.base_open, - c_iso := λ U, - begin - generalize_proofs h, - dsimp only [algebraic_geometry.PresheafedSpace.comp_c_app, unop_op, functor.op, comp_base, - Top.presheaf.pushforward_obj_obj, opens.map_comp_obj], - apply_with is_iso.comp_is_iso { instances := ff }, - swap, - { have : (opens.map g.base).obj (h.functor.obj U) = hf.open_functor.obj U, - { ext1, - dsimp only [opens.map_coe, is_open_map.functor_obj_coe, comp_base], - rw [coe_comp, ← set.image_image, set.preimage_image_eq _ hg.base_open.inj] }, - rw this, - apply_instance }, - { have : h.functor.obj U = hg.open_functor.obj (hf.open_functor.obj U), - { ext1, - dsimp only [is_open_map.functor_obj_coe], - rw [comp_base, coe_comp, ←set.image_image] }, - rw this, - apply_instance } - end } - -/-- For an open immersion `f : X ⟶ Y` and an open set `U ⊆ X`, we have the map `X(U) ⟶ Y(U)`. -/ -noncomputable -def inv_app (U : opens X) : X.presheaf.obj (op U) ⟶ Y.presheaf.obj (op (H.open_functor.obj U)) := -X.presheaf.map (eq_to_hom (by simp [opens.map, set.preimage_image_eq _ H.base_open.inj])) ≫ - inv (f.c.app (op (H.open_functor.obj U))) - -@[simp, reassoc] lemma inv_naturality {U V : (opens X)ᵒᵖ} (i : U ⟶ V) : - X.presheaf.map i ≫ H.inv_app (unop V) = H.inv_app (unop U) ≫ - Y.presheaf.map (H.open_functor.op.map i) := -begin - simp only [inv_app, ←category.assoc], - rw [is_iso.comp_inv_eq], - simp only [category.assoc, f.c.naturality, is_iso.inv_hom_id_assoc, ← X.presheaf.map_comp], - erw ← X.presheaf.map_comp, - congr -end - -instance (U : opens X) : is_iso (H.inv_app U) := by { delta inv_app, apply_instance } - -lemma inv_inv_app (U : opens X) : - inv (H.inv_app U) = f.c.app (op (H.open_functor.obj U)) ≫ - X.presheaf.map (eq_to_hom (by simp [opens.map, set.preimage_image_eq _ H.base_open.inj])) := -begin - rw ← cancel_epi (H.inv_app U), - rw is_iso.hom_inv_id, - delta inv_app, - simp [← functor.map_comp] -end - -@[simp, reassoc, elementwise] lemma inv_app_app (U : opens X) : - H.inv_app U ≫ f.c.app (op (H.open_functor.obj U)) = - X.presheaf.map (eq_to_hom (by simp [opens.map, set.preimage_image_eq _ H.base_open.inj])) := -by rw [inv_app, category.assoc, is_iso.inv_hom_id, category.comp_id] - -@[simp, reassoc] lemma app_inv_app (U : opens Y) : - f.c.app (op U) ≫ H.inv_app ((opens.map f.base).obj U) = - Y.presheaf.map ((hom_of_le (by exact set.image_preimage_subset f.base U)).op : - op U ⟶ op (H.open_functor.obj ((opens.map f.base).obj U))) := -by { erw ← category.assoc, rw [is_iso.comp_inv_eq, f.c.naturality], congr } - -/-- A variant of `app_inv_app` that gives an `eq_to_hom` instead of `hom_of_le`. -/ -@[reassoc] lemma app_inv_app' (U : opens Y) (hU : (U : set Y) ⊆ set.range f.base) : - f.c.app (op U) ≫ H.inv_app ((opens.map f.base).obj U) = - Y.presheaf.map (eq_to_hom (by - { apply le_antisymm, - { exact set.image_preimage_subset f.base U.1 }, - { rw [← set_like.coe_subset_coe], - refine has_le.le.trans_eq _ (@set.image_preimage_eq_inter_range _ _ f.base U.1).symm, - exact set.subset_inter_iff.mpr ⟨λ _ h, h, hU⟩ } })).op := -by { erw ← category.assoc, rw [is_iso.comp_inv_eq, f.c.naturality], congr } - -/-- An isomorphism is an open immersion. -/ -instance of_iso {X Y : PresheafedSpace.{v} C} (H : X ≅ Y) : is_open_immersion H.hom := -{ base_open := (Top.homeo_of_iso ((forget C).map_iso H)).open_embedding, - c_iso := λ _, infer_instance } - -@[priority 100] -instance of_is_iso {X Y : PresheafedSpace.{v} C} (f : X ⟶ Y) [is_iso f] : is_open_immersion f := -algebraic_geometry.PresheafedSpace.is_open_immersion.of_iso (as_iso f) - -instance of_restrict {X : Top} (Y : PresheafedSpace C) {f : X ⟶ Y.carrier} - (hf : open_embedding f) : is_open_immersion (Y.of_restrict hf) := -{ base_open := hf, - c_iso := λ U, - begin - dsimp, - have : (opens.map f).obj (hf.is_open_map.functor.obj U) = U, - { ext1, - exact set.preimage_image_eq _ hf.inj }, - convert (show is_iso (Y.presheaf.map (𝟙 _)), from infer_instance), - { apply subsingleton.helim, - rw this }, - { rw Y.presheaf.map_id, - apply_instance } - end } - -@[elementwise, simp] -lemma of_restrict_inv_app {C : Type*} [category C] (X : PresheafedSpace C) {Y : Top} - {f : Y ⟶ Top.of X.carrier} - (h : open_embedding f) (U : opens (X.restrict h).carrier) : - (PresheafedSpace.is_open_immersion.of_restrict X h).inv_app U = 𝟙 _ := -begin - delta PresheafedSpace.is_open_immersion.inv_app, - rw [is_iso.comp_inv_eq, category.id_comp], - change X.presheaf.map _ = X.presheaf.map _, - congr -end - -/-- An open immersion is an iso if the underlying continuous map is epi. -/ -lemma to_iso (f : X ⟶ Y) [h : is_open_immersion f] [h' : epi f.base] : is_iso f := -begin - apply_with is_iso_of_components { instances := ff }, - { let : X ≃ₜ Y := (homeomorph.of_embedding _ h.base_open.to_embedding).trans - { to_fun := subtype.val, inv_fun := λ x, ⟨x, - by { rw set.range_iff_surjective.mpr ((Top.epi_iff_surjective _).mp h'), trivial }⟩, - left_inv := λ ⟨_,_⟩, rfl, right_inv := λ _, rfl }, - convert is_iso.of_iso (Top.iso_of_homeo this), - { ext, refl } }, - { apply_with nat_iso.is_iso_of_is_iso_app { instances := ff }, - intro U, - have : U = op (h.open_functor.obj ((opens.map f.base).obj (unop U))), - { induction U using opposite.rec, - cases U, - dsimp only [functor.op, opens.map], - congr, - exact (set.image_preimage_eq _ ((Top.epi_iff_surjective _).mp h')).symm }, - convert @@is_open_immersion.c_iso _ h ((opens.map f.base).obj (unop U)) } -end - -instance stalk_iso [has_colimits C] [H : is_open_immersion f] (x : X) : is_iso (stalk_map f x) := -begin - rw ← H.iso_restrict_hom_of_restrict, - rw PresheafedSpace.stalk_map.comp, - apply_instance -end - -end - -section pullback - -noncomputable theory - -variables {X Y Z : PresheafedSpace.{v} C} (f : X ⟶ Z) [hf : is_open_immersion f] (g : Y ⟶ Z) - -include hf - -/-- - (Implementation.) The projection map when constructing the pullback along an open immersion. --/ -def pullback_cone_of_left_fst : - Y.restrict (Top.snd_open_embedding_of_left_open_embedding hf.base_open g.base) ⟶ X := -{ base := pullback.fst, - c := - { app := λ U, hf.inv_app (unop U) ≫ - g.c.app (op (hf.base_open.is_open_map.functor.obj (unop U))) ≫ - Y.presheaf.map (eq_to_hom - (begin - simp only [is_open_map.functor, subtype.mk_eq_mk, unop_op, op_inj_iff, opens.map, - subtype.coe_mk, functor.op_obj, subtype.val_eq_coe], - apply has_le.le.antisymm, - { rintros _ ⟨_, h₁, h₂⟩, - use (Top.pullback_iso_prod_subtype _ _).inv ⟨⟨_, _⟩, h₂⟩, - simpa using h₁ }, - { rintros _ ⟨x, h₁, rfl⟩, - exact ⟨_, h₁, concrete_category.congr_hom pullback.condition x⟩ } - end)), - naturality' := - begin - intros U V i, - induction U using opposite.rec, - induction V using opposite.rec, - simp only [quiver.hom.unop_op, Top.presheaf.pushforward_obj_map, category.assoc, - nat_trans.naturality_assoc, functor.op_map, inv_naturality_assoc, ← Y.presheaf.map_comp], - erw ← Y.presheaf.map_comp, - congr - end } } - -lemma pullback_cone_of_left_condition : - pullback_cone_of_left_fst f g ≫ f = Y.of_restrict _ ≫ g := -begin - ext U, - { induction U using opposite.rec, - dsimp only [comp_c_app, nat_trans.comp_app, unop_op, - whisker_right_app, pullback_cone_of_left_fst], - simp only [quiver.hom.unop_op, Top.presheaf.pushforward_obj_map, app_inv_app_assoc, - eq_to_hom_app, eq_to_hom_unop, category.assoc, nat_trans.naturality_assoc, functor.op_map], - erw [← Y.presheaf.map_comp, ← Y.presheaf.map_comp], - congr }, - { simpa using pullback.condition } -end - -/-- -We construct the pullback along an open immersion via restricting along the pullback of the -maps of underlying spaces (which is also an open embedding). --/ -def pullback_cone_of_left : pullback_cone f g := -pullback_cone.mk (pullback_cone_of_left_fst f g) (Y.of_restrict _) - (pullback_cone_of_left_condition f g) - -variable (s : pullback_cone f g) - -/-- - (Implementation.) Any cone over `cospan f g` indeed factors through the constructed cone. --/ -def pullback_cone_of_left_lift : s.X ⟶ (pullback_cone_of_left f g).X := -{ base := pullback.lift s.fst.base s.snd.base - (congr_arg (λ x, PresheafedSpace.hom.base x) s.condition), - c := - { app := λ U, s.snd.c.app _ ≫ s.X.presheaf.map (eq_to_hom (begin - dsimp only [opens.map, is_open_map.functor, functor.op], - congr' 2, - let s' : pullback_cone f.base g.base := pullback_cone.mk s.fst.base s.snd.base _, - have : _ = s.snd.base := limit.lift_π s' walking_cospan.right, - conv_lhs { erw ← this, rw coe_comp, erw ← set.preimage_preimage }, - erw set.preimage_image_eq _ - (Top.snd_open_embedding_of_left_open_embedding hf.base_open g.base).inj, - end)), - naturality' := λ U V i, - begin - erw s.snd.c.naturality_assoc, - rw category.assoc, - erw [← s.X.presheaf.map_comp, ← s.X.presheaf.map_comp], - congr - end } } - --- this lemma is not a `simp` lemma, because it is an implementation detail -lemma pullback_cone_of_left_lift_fst : - pullback_cone_of_left_lift f g s ≫ (pullback_cone_of_left f g).fst = s.fst := -begin - ext x, - { induction x using opposite.rec, - change ((_ ≫ _) ≫ _ ≫ _) ≫ _ = _, - simp_rw [category.assoc], - erw ← s.X.presheaf.map_comp, - erw s.snd.c.naturality_assoc, - have := congr_app s.condition (op (hf.open_functor.obj x)), - dsimp only [comp_c_app, unop_op] at this, - rw ← is_iso.comp_inv_eq at this, - reassoc! this, - erw [← this, hf.inv_app_app_assoc, s.fst.c.naturality_assoc], - simpa [eq_to_hom_map], }, - { change pullback.lift _ _ _ ≫ pullback.fst = _, - simp } -end - --- this lemma is not a `simp` lemma, because it is an implementation detail -lemma pullback_cone_of_left_lift_snd : - pullback_cone_of_left_lift f g s ≫ (pullback_cone_of_left f g).snd = s.snd := -begin - ext x, - { change (_ ≫ _ ≫ _) ≫ _ = _, - simp_rw category.assoc, - erw s.snd.c.naturality_assoc, - erw [← s.X.presheaf.map_comp, ← s.X.presheaf.map_comp], - transitivity s.snd.c.app x ≫ s.X.presheaf.map (𝟙 _), - { congr }, - { rw s.X.presheaf.map_id, erw category.comp_id } }, - { change pullback.lift _ _ _ ≫ pullback.snd = _, - simp } -end - -instance pullback_cone_snd_is_open_immersion : - is_open_immersion (pullback_cone_of_left f g).snd := -begin - erw category_theory.limits.pullback_cone.mk_snd, - apply_instance -end - -/-- The constructed pullback cone is indeed the pullback. -/ -def pullback_cone_of_left_is_limit : - is_limit (pullback_cone_of_left f g) := -begin - apply pullback_cone.is_limit_aux', - intro s, - use pullback_cone_of_left_lift f g s, - use pullback_cone_of_left_lift_fst f g s, - use pullback_cone_of_left_lift_snd f g s, - intros m h₁ h₂, - rw ← cancel_mono (pullback_cone_of_left f g).snd, - exact (h₂.trans (pullback_cone_of_left_lift_snd f g s).symm) -end - -instance has_pullback_of_left : - has_pullback f g := -⟨⟨⟨_, pullback_cone_of_left_is_limit f g⟩⟩⟩ - -instance has_pullback_of_right : - has_pullback g f := has_pullback_symmetry f g - -/-- Open immersions are stable under base-change. -/ -instance pullback_snd_of_left : - is_open_immersion (pullback.snd : pullback f g ⟶ _) := -begin - delta pullback.snd, - rw ← limit.iso_limit_cone_hom_π ⟨_, pullback_cone_of_left_is_limit f g⟩ walking_cospan.right, - apply_instance -end - -/-- Open immersions are stable under base-change. -/ -instance pullback_fst_of_right : - is_open_immersion (pullback.fst : pullback g f ⟶ _) := -begin - rw ← pullback_symmetry_hom_comp_snd, - apply_instance -end - -instance pullback_to_base_is_open_immersion [is_open_immersion g] : - is_open_immersion (limit.π (cospan f g) walking_cospan.one) := -begin - rw [←limit.w (cospan f g) walking_cospan.hom.inl, cospan_map_inl], - apply_instance -end - -instance forget_preserves_limits_of_left : preserves_limit (cospan f g) (forget C) := -preserves_limit_of_preserves_limit_cone (pullback_cone_of_left_is_limit f g) -begin - apply (is_limit.postcompose_hom_equiv (diagram_iso_cospan.{v} _) _).to_fun, - refine (is_limit.equiv_iso_limit _).to_fun (limit.is_limit (cospan f.base g.base)), - fapply cones.ext, - exact (iso.refl _), - change ∀ j, _ = 𝟙 _ ≫ _ ≫ _, - simp_rw category.id_comp, - rintros (_|_|_); symmetry, - { erw category.comp_id, - exact limit.w (cospan f.base g.base) walking_cospan.hom.inl }, - { exact category.comp_id _ }, - { exact category.comp_id _ }, -end - -instance forget_preserves_limits_of_right : preserves_limit (cospan g f) (forget C) := -preserves_pullback_symmetry (forget C) f g - -lemma pullback_snd_is_iso_of_range_subset (H : set.range g.base ⊆ set.range f.base) : - is_iso (pullback.snd : pullback f g ⟶ _) := -begin - haveI := Top.snd_iso_of_left_embedding_range_subset hf.base_open.to_embedding g.base H, - haveI : is_iso (pullback.snd : pullback f g ⟶ _).base, - { delta pullback.snd, - rw ← limit.iso_limit_cone_hom_π ⟨_, pullback_cone_of_left_is_limit f g⟩ walking_cospan.right, - change is_iso (_ ≫ pullback.snd), - apply_instance }, - apply to_iso -end - -/-- -The universal property of open immersions: -For an open immersion `f : X ⟶ Z`, given any morphism of schemes `g : Y ⟶ Z` whose topological -image is contained in the image of `f`, we can lift this morphism to a unique `Y ⟶ X` that -commutes with these maps. --/ -def lift (H : set.range g.base ⊆ set.range f.base) : Y ⟶ X := -begin - haveI := pullback_snd_is_iso_of_range_subset f g H, - exact inv (pullback.snd : pullback f g ⟶ _) ≫ pullback.fst, -end - -@[simp, reassoc] lemma lift_fac (H : set.range g.base ⊆ set.range f.base) : - lift f g H ≫ f = g := -by { erw category.assoc, rw is_iso.inv_comp_eq, exact pullback.condition } - -lemma lift_uniq (H : set.range g.base ⊆ set.range f.base) (l : Y ⟶ X) - (hl : l ≫ f = g) : l = lift f g H := -by rw [← cancel_mono f, hl, lift_fac] - -/-- Two open immersions with equal range is isomorphic. -/ -@[simps] def iso_of_range_eq [is_open_immersion g] (e : set.range f.base = set.range g.base) : - X ≅ Y := -{ hom := lift g f (le_of_eq e), - inv := lift f g (le_of_eq e.symm), - hom_inv_id' := by { rw ← cancel_mono f, simp }, - inv_hom_id' := by { rw ← cancel_mono g, simp } } - -end pullback - -open category_theory.limits.walking_cospan - -section to_SheafedSpace - -variables {X : PresheafedSpace.{v} C} (Y : SheafedSpace C) -variables (f : X ⟶ Y.to_PresheafedSpace) [H : is_open_immersion f] - -include H - -/-- If `X ⟶ Y` is an open immersion, and `Y` is a SheafedSpace, then so is `X`. -/ -def to_SheafedSpace : SheafedSpace C := -{ is_sheaf := - begin - apply Top.presheaf.is_sheaf_of_iso (sheaf_iso_of_iso H.iso_restrict.symm).symm, - apply Top.sheaf.pushforward_sheaf_of_sheaf, - exact (Y.restrict H.base_open).is_sheaf - end, - to_PresheafedSpace := X } - -@[simp] lemma to_SheafedSpace_to_PresheafedSpace : (to_SheafedSpace Y f).to_PresheafedSpace = X := -rfl - -/-- -If `X ⟶ Y` is an open immersion of PresheafedSpaces, and `Y` is a SheafedSpace, we can -upgrade it into a morphism of SheafedSpaces. --/ -def to_SheafedSpace_hom : to_SheafedSpace Y f ⟶ Y := f - -@[simp] lemma to_SheafedSpace_hom_base : (to_SheafedSpace_hom Y f).base = f.base := rfl - -@[simp] lemma to_SheafedSpace_hom_c : (to_SheafedSpace_hom Y f).c = f.c := rfl - -instance to_SheafedSpace_is_open_immersion : - SheafedSpace.is_open_immersion (to_SheafedSpace_hom Y f) := H - -omit H - -@[simp] lemma SheafedSpace_to_SheafedSpace {X Y : SheafedSpace.{v} C} (f : X ⟶ Y) - [is_open_immersion f] : to_SheafedSpace Y f = X := by unfreezingI { cases X, refl } - -end to_SheafedSpace - -section to_LocallyRingedSpace - -variables {X : PresheafedSpace.{u} CommRing.{u}} (Y : LocallyRingedSpace.{u}) -variables (f : X ⟶ Y.to_PresheafedSpace) [H : is_open_immersion f] - -include H - -/-- If `X ⟶ Y` is an open immersion, and `Y` is a LocallyRingedSpace, then so is `X`. -/ -def to_LocallyRingedSpace : LocallyRingedSpace := -{ to_SheafedSpace := to_SheafedSpace Y.to_SheafedSpace f, - local_ring := λ x, begin - haveI : local_ring (Y.to_SheafedSpace.to_PresheafedSpace.stalk (f.base x)) := Y.local_ring _, - exact (as_iso (stalk_map f x)).CommRing_iso_to_ring_equiv.local_ring - end } - -@[simp] lemma to_LocallyRingedSpace_to_SheafedSpace : - (to_LocallyRingedSpace Y f).to_SheafedSpace = (to_SheafedSpace Y.1 f) := rfl - -/-- -If `X ⟶ Y` is an open immersion of PresheafedSpaces, and `Y` is a LocallyRingedSpace, we can -upgrade it into a morphism of LocallyRingedSpace. --/ -def to_LocallyRingedSpace_hom : to_LocallyRingedSpace Y f ⟶ Y := ⟨f, λ x, infer_instance⟩ - -@[simp] lemma to_LocallyRingedSpace_hom_val : - (to_LocallyRingedSpace_hom Y f).val = f := rfl - -instance to_LocallyRingedSpace_is_open_immersion : - LocallyRingedSpace.is_open_immersion (to_LocallyRingedSpace_hom Y f) := H - -omit H - -@[simp] lemma LocallyRingedSpace_to_LocallyRingedSpace {X Y : LocallyRingedSpace} (f : X ⟶ Y) - [LocallyRingedSpace.is_open_immersion f] : - to_LocallyRingedSpace Y f.1 = X := -by unfreezingI { cases X, delta to_LocallyRingedSpace, simp } - -end to_LocallyRingedSpace - -lemma is_iso_of_subset {X Y : PresheafedSpace.{v} C} (f : X ⟶ Y) - [H : PresheafedSpace.is_open_immersion f] (U : opens Y.carrier) - (hU : (U : set Y.carrier) ⊆ set.range f.base) : is_iso (f.c.app $ op U) := -begin - have : U = H.base_open.is_open_map.functor.obj ((opens.map f.base).obj U), - { ext1, - exact (set.inter_eq_left_iff_subset.mpr hU).symm.trans set.image_preimage_eq_inter_range.symm }, - convert PresheafedSpace.is_open_immersion.c_iso ((opens.map f.base).obj U), -end - -end PresheafedSpace.is_open_immersion - -namespace SheafedSpace.is_open_immersion - -@[priority 100] -instance of_is_iso {X Y : SheafedSpace.{v} C} (f : X ⟶ Y) [is_iso f] : - SheafedSpace.is_open_immersion f := -@@PresheafedSpace.is_open_immersion.of_is_iso _ f -(SheafedSpace.forget_to_PresheafedSpace.map_is_iso _) - -instance comp {X Y Z : SheafedSpace C} (f : X ⟶ Y) (g : Y ⟶ Z) - [SheafedSpace.is_open_immersion f] [SheafedSpace.is_open_immersion g] : - SheafedSpace.is_open_immersion (f ≫ g) := PresheafedSpace.is_open_immersion.comp f g - -section pullback - -variables {X Y Z : SheafedSpace C} (f : X ⟶ Z) (g : Y ⟶ Z) -variable [H : SheafedSpace.is_open_immersion f] - -include H - -local notation `forget` := SheafedSpace.forget_to_PresheafedSpace -open category_theory.limits.walking_cospan - -instance : mono f := -forget .mono_of_mono_map (show @mono (PresheafedSpace C) _ _ _ f, by apply_instance) - -instance forget_map_is_open_immersion : - PresheafedSpace.is_open_immersion (forget .map f) := ⟨H.base_open, H.c_iso⟩ - -instance has_limit_cospan_forget_of_left : has_limit (cospan f g ⋙ forget) := -begin - apply has_limit_of_iso (diagram_iso_cospan.{v} _).symm, - change has_limit (cospan (forget .map f) (forget .map g)), - apply_instance -end - -instance has_limit_cospan_forget_of_left' : has_limit (cospan ((cospan f g ⋙ forget).map hom.inl) - ((cospan f g ⋙ forget).map hom.inr)) := -show has_limit (cospan (forget .map f) (forget .map g)), from infer_instance - -instance has_limit_cospan_forget_of_right : has_limit (cospan g f ⋙ forget) := -begin - apply has_limit_of_iso (diagram_iso_cospan.{v} _).symm, - change has_limit (cospan (forget .map g) (forget .map f)), - apply_instance -end - -instance has_limit_cospan_forget_of_right' : has_limit (cospan ((cospan g f ⋙ forget).map hom.inl) - ((cospan g f ⋙ forget).map hom.inr)) := -show has_limit (cospan (forget .map g) (forget .map f)), from infer_instance - - -instance forget_creates_pullback_of_left : creates_limit (cospan f g) forget := -creates_limit_of_fully_faithful_of_iso - (PresheafedSpace.is_open_immersion.to_SheafedSpace Y - (@pullback.snd (PresheafedSpace C) _ _ _ _ f g _)) - (eq_to_iso (show pullback _ _ = pullback _ _, by congr) - ≪≫ has_limit.iso_of_nat_iso (diagram_iso_cospan _).symm) - -instance forget_creates_pullback_of_right : creates_limit (cospan g f) forget := -creates_limit_of_fully_faithful_of_iso - (PresheafedSpace.is_open_immersion.to_SheafedSpace Y - (@pullback.fst (PresheafedSpace C) _ _ _ _ g f _)) - (eq_to_iso (show pullback _ _ = pullback _ _, by congr) - ≪≫ has_limit.iso_of_nat_iso (diagram_iso_cospan _).symm) - -instance SheafedSpace_forget_preserves_of_left : - preserves_limit (cospan f g) (SheafedSpace.forget C) := -@@limits.comp_preserves_limit _ _ _ _ forget (PresheafedSpace.forget C) _ -begin - apply_with (preserves_limit_of_iso_diagram _ (diagram_iso_cospan.{v} _).symm) { instances := tt }, - dsimp, - apply_instance -end - -instance SheafedSpace_forget_preserves_of_right : - preserves_limit (cospan g f) (SheafedSpace.forget C) := -preserves_pullback_symmetry _ _ _ - -instance SheafedSpace_has_pullback_of_left : has_pullback f g := - has_limit_of_created (cospan f g) forget - -instance SheafedSpace_has_pullback_of_right : has_pullback g f := - has_limit_of_created (cospan g f) forget - -/-- Open immersions are stable under base-change. -/ -instance SheafedSpace_pullback_snd_of_left : - SheafedSpace.is_open_immersion (pullback.snd : pullback f g ⟶ _) := -begin - delta pullback.snd, - have : _ = limit.π (cospan f g) right := preserves_limits_iso_hom_π - forget (cospan f g) right, - rw ← this, - have := has_limit.iso_of_nat_iso_hom_π - (diagram_iso_cospan.{v} (cospan f g ⋙ forget)) - right, - erw category.comp_id at this, - rw ← this, - dsimp, - apply_instance -end - -instance SheafedSpace_pullback_fst_of_right : - SheafedSpace.is_open_immersion (pullback.fst : pullback g f ⟶ _) := -begin - delta pullback.fst, - have : _ = limit.π (cospan g f) left := preserves_limits_iso_hom_π - forget (cospan g f) left, - rw ← this, - have := has_limit.iso_of_nat_iso_hom_π - (diagram_iso_cospan.{v} (cospan g f ⋙ forget)) left, - erw category.comp_id at this, - rw ← this, - dsimp, - apply_instance -end - -instance SheafedSpace_pullback_to_base_is_open_immersion [SheafedSpace.is_open_immersion g] : - SheafedSpace.is_open_immersion (limit.π (cospan f g) one : pullback f g ⟶ Z) := -begin - rw [←limit.w (cospan f g) hom.inl, cospan_map_inl], - apply_instance -end - -end pullback - -section of_stalk_iso -variables [has_limits C] [has_colimits C] [concrete_category.{v} C] -variables [reflects_isomorphisms (forget C)] [preserves_limits (forget C)] -variables [preserves_filtered_colimits (forget C)] - -/-- -Suppose `X Y : SheafedSpace C`, where `C` is a concrete category, -whose forgetful functor reflects isomorphisms, preserves limits and filtered colimits. -Then a morphism `X ⟶ Y` that is a topological open embedding -is an open immersion iff every stalk map is an iso. --/ -lemma of_stalk_iso {X Y : SheafedSpace C} (f : X ⟶ Y) - (hf : open_embedding f.base) [H : ∀ x : X, is_iso (PresheafedSpace.stalk_map f x)] : - SheafedSpace.is_open_immersion f := -{ base_open := hf, - c_iso := λ U, begin - apply_with (Top.presheaf.app_is_iso_of_stalk_functor_map_iso - (show Y.sheaf ⟶ (Top.sheaf.pushforward f.base).obj X.sheaf, from ⟨f.c⟩)) { instances := ff }, - rintros ⟨_, y, hy, rfl⟩, - specialize H y, - delta PresheafedSpace.stalk_map at H, - haveI H' := Top.presheaf.stalk_pushforward.stalk_pushforward_iso_of_open_embedding - C hf X.presheaf y, - have := @@is_iso.comp_is_iso _ H (@@is_iso.inv_is_iso _ H'), - rw [category.assoc, is_iso.hom_inv_id, category.comp_id] at this, - exact this - end } - -end of_stalk_iso - -section prod - -variables [has_limits C] {ι : Type v} (F : discrete ι ⥤ SheafedSpace C) [has_colimit F] - (i : discrete ι) - -lemma sigma_ι_open_embedding : open_embedding (colimit.ι F i).base := -begin - rw ← (show _ = (colimit.ι F i).base, - from ι_preserves_colimits_iso_inv (SheafedSpace.forget C) F i), - have : _ = _ ≫ colimit.ι (discrete.functor ((F ⋙ SheafedSpace.forget C).obj ∘ discrete.mk)) i := - has_colimit.iso_of_nat_iso_ι_hom discrete.nat_iso_functor i, - rw ← iso.eq_comp_inv at this, - rw this, - have : colimit.ι _ _ ≫ _ = _ := - Top.sigma_iso_sigma_hom_ι.{v v} ((F ⋙ SheafedSpace.forget C).obj ∘ discrete.mk) i.as, - rw ← iso.eq_comp_inv at this, - cases i, - rw this, - simp_rw [← category.assoc, Top.open_embedding_iff_comp_is_iso, - Top.open_embedding_iff_is_iso_comp], - dsimp, - exact open_embedding_sigma_mk -end - -lemma image_preimage_is_empty (j : discrete ι) (h : i ≠ j) (U : opens (F.obj i)) : - (opens.map (colimit.ι (F ⋙ SheafedSpace.forget_to_PresheafedSpace) j).base).obj - ((opens.map (preserves_colimit_iso SheafedSpace.forget_to_PresheafedSpace F).inv.base).obj - ((sigma_ι_open_embedding F i).is_open_map.functor.obj U)) = ⊥ := -begin - ext, - apply iff_false_intro, - rintro ⟨y, hy, eq⟩, - replace eq := concrete_category.congr_arg - (preserves_colimit_iso (SheafedSpace.forget C) F ≪≫ - has_colimit.iso_of_nat_iso discrete.nat_iso_functor ≪≫ Top.sigma_iso_sigma.{v} _).hom eq, - simp_rw [category_theory.iso.trans_hom, ← Top.comp_app, ← PresheafedSpace.comp_base] at eq, - rw ι_preserves_colimits_iso_inv at eq, - change ((SheafedSpace.forget C).map (colimit.ι F i) ≫ _) y = - ((SheafedSpace.forget C).map (colimit.ι F j) ≫ _) x at eq, - cases i, cases j, - rw [ι_preserves_colimits_iso_hom_assoc, ι_preserves_colimits_iso_hom_assoc, - has_colimit.iso_of_nat_iso_ι_hom_assoc, has_colimit.iso_of_nat_iso_ι_hom_assoc, - Top.sigma_iso_sigma_hom_ι.{v}, Top.sigma_iso_sigma_hom_ι.{v}] at eq, - exact h (congr_arg discrete.mk (congr_arg sigma.fst eq)), -end - -instance sigma_ι_is_open_immersion [has_strict_terminal_objects C] : - SheafedSpace.is_open_immersion (colimit.ι F i) := -{ base_open := sigma_ι_open_embedding F i, - c_iso := λ U, begin - have e : colimit.ι F i = _ := - (ι_preserves_colimits_iso_inv SheafedSpace.forget_to_PresheafedSpace F i).symm, - have H : open_embedding (colimit.ι (F ⋙ SheafedSpace.forget_to_PresheafedSpace) i ≫ - (preserves_colimit_iso SheafedSpace.forget_to_PresheafedSpace F).inv).base := - e ▸ sigma_ι_open_embedding F i, - suffices : is_iso ((colimit.ι (F ⋙ SheafedSpace.forget_to_PresheafedSpace) i ≫ - (preserves_colimit_iso SheafedSpace.forget_to_PresheafedSpace F).inv).c.app - (op (H.is_open_map.functor.obj U))), - { convert this }, - rw [PresheafedSpace.comp_c_app, - ← PresheafedSpace.colimit_presheaf_obj_iso_componentwise_limit_hom_π], - rsufficesI : is_iso (limit.π (PresheafedSpace.componentwise_diagram - (F ⋙ SheafedSpace.forget_to_PresheafedSpace) - ((opens.map (preserves_colimit_iso SheafedSpace.forget_to_PresheafedSpace F).inv.base).obj - (unop $ op $ H.is_open_map.functor.obj U))) (op i)), - { apply_instance }, - apply limit_π_is_iso_of_is_strict_terminal, - intros j hj, - induction j using opposite.rec, - dsimp, - convert (F.obj j).sheaf.is_terminal_of_empty, - convert image_preimage_is_empty F i j (λ h, hj (congr_arg op h.symm)) U, - exact (congr_arg PresheafedSpace.hom.base e).symm - end } - -end prod - -end SheafedSpace.is_open_immersion - namespace LocallyRingedSpace.is_open_immersion -section pullback - -variables {X Y Z : LocallyRingedSpace.{u}} (f : X ⟶ Z) (g : Y ⟶ Z) -variable [H : LocallyRingedSpace.is_open_immersion f] - -@[priority 100] -instance of_is_iso [is_iso g] : - LocallyRingedSpace.is_open_immersion g := -@@PresheafedSpace.is_open_immersion.of_is_iso _ g.1 ⟨⟨(inv g).1, - by { erw ← LocallyRingedSpace.comp_val, rw is_iso.hom_inv_id, - erw ← LocallyRingedSpace.comp_val, rw is_iso.inv_hom_id, split; simpa }⟩⟩ - -include H - -instance comp (g : Z ⟶ Y) [LocallyRingedSpace.is_open_immersion g] : - LocallyRingedSpace.is_open_immersion (f ≫ g) := PresheafedSpace.is_open_immersion.comp f.1 g.1 - -instance mono : mono f := -LocallyRingedSpace.forget_to_SheafedSpace.mono_of_mono_map (show mono f.1, by apply_instance) - -instance : SheafedSpace.is_open_immersion (LocallyRingedSpace.forget_to_SheafedSpace.map f) := H - -/-- An explicit pullback cone over `cospan f g` if `f` is an open immersion. -/ -def pullback_cone_of_left : pullback_cone f g := -begin - refine pullback_cone.mk _ - (Y.of_restrict (Top.snd_open_embedding_of_left_open_embedding H.base_open g.1.base)) _, - { use PresheafedSpace.is_open_immersion.pullback_cone_of_left_fst f.1 g.1, - intro x, - have := PresheafedSpace.stalk_map.congr_hom _ _ - (PresheafedSpace.is_open_immersion.pullback_cone_of_left_condition f.1 g.1) x, - rw [PresheafedSpace.stalk_map.comp, PresheafedSpace.stalk_map.comp] at this, - rw ← is_iso.eq_inv_comp at this, - rw this, - apply_instance }, - { exact LocallyRingedSpace.hom.ext _ _ - (PresheafedSpace.is_open_immersion.pullback_cone_of_left_condition _ _) }, -end - -instance : LocallyRingedSpace.is_open_immersion (pullback_cone_of_left f g).snd := -show PresheafedSpace.is_open_immersion (Y.to_PresheafedSpace.of_restrict _), by apply_instance - -/-- The constructed `pullback_cone_of_left` is indeed limiting. -/ -def pullback_cone_of_left_is_limit : is_limit (pullback_cone_of_left f g) := -pullback_cone.is_limit_aux' _ $ λ s, -begin - use PresheafedSpace.is_open_immersion.pullback_cone_of_left_lift f.1 g.1 - (pullback_cone.mk s.fst.1 s.snd.1 (congr_arg LocallyRingedSpace.hom.val s.condition)), - { intro x, - have := PresheafedSpace.stalk_map.congr_hom _ _ - (PresheafedSpace.is_open_immersion.pullback_cone_of_left_lift_snd f.1 g.1 - (pullback_cone.mk s.fst.1 s.snd.1 (congr_arg LocallyRingedSpace.hom.val s.condition))) x, - change _ = _ ≫ PresheafedSpace.stalk_map s.snd.1 x at this, - rw [PresheafedSpace.stalk_map.comp, ← is_iso.eq_inv_comp] at this, - rw this, - apply_instance }, - split, - { exact LocallyRingedSpace.hom.ext _ _ - (PresheafedSpace.is_open_immersion.pullback_cone_of_left_lift_fst f.1 g.1 _) }, - split, - { exact LocallyRingedSpace.hom.ext _ _ - (PresheafedSpace.is_open_immersion.pullback_cone_of_left_lift_snd f.1 g.1 _) }, - intros m h₁ h₂, - rw ← cancel_mono (pullback_cone_of_left f g).snd, - exact (h₂.trans (LocallyRingedSpace.hom.ext _ _ - (PresheafedSpace.is_open_immersion.pullback_cone_of_left_lift_snd f.1 g.1 - (pullback_cone.mk s.fst.1 s.snd.1 (congr_arg LocallyRingedSpace.hom.val s.condition))).symm)) -end - -instance has_pullback_of_left : - has_pullback f g := -⟨⟨⟨_, pullback_cone_of_left_is_limit f g⟩⟩⟩ - -instance has_pullback_of_right : - has_pullback g f := has_pullback_symmetry f g - -/-- Open immersions are stable under base-change. -/ -instance pullback_snd_of_left : - LocallyRingedSpace.is_open_immersion (pullback.snd : pullback f g ⟶ _) := -begin - delta pullback.snd, - rw ← limit.iso_limit_cone_hom_π ⟨_, pullback_cone_of_left_is_limit f g⟩ walking_cospan.right, - apply_instance -end - -/-- Open immersions are stable under base-change. -/ -instance pullback_fst_of_right : -LocallyRingedSpace.is_open_immersion (pullback.fst : pullback g f ⟶ _) := -begin - rw ← pullback_symmetry_hom_comp_snd, - apply_instance -end - -instance pullback_to_base_is_open_immersion [LocallyRingedSpace.is_open_immersion g] : - LocallyRingedSpace.is_open_immersion (limit.π (cospan f g) walking_cospan.one) := -begin - rw [←limit.w (cospan f g) walking_cospan.hom.inl, cospan_map_inl], - apply_instance -end - -instance forget_preserves_pullback_of_left : - preserves_limit (cospan f g) LocallyRingedSpace.forget_to_SheafedSpace := -preserves_limit_of_preserves_limit_cone (pullback_cone_of_left_is_limit f g) -begin - apply (is_limit_map_cone_pullback_cone_equiv _ _).symm.to_fun, - apply is_limit_of_is_limit_pullback_cone_map SheafedSpace.forget_to_PresheafedSpace, - exact PresheafedSpace.is_open_immersion.pullback_cone_of_left_is_limit f.1 g.1 -end - -instance forget_to_PresheafedSpace_preserves_pullback_of_left : - preserves_limit (cospan f g) - (LocallyRingedSpace.forget_to_SheafedSpace ⋙ SheafedSpace.forget_to_PresheafedSpace) := -preserves_limit_of_preserves_limit_cone (pullback_cone_of_left_is_limit f g) -begin - apply (is_limit_map_cone_pullback_cone_equiv _ _).symm.to_fun, - exact PresheafedSpace.is_open_immersion.pullback_cone_of_left_is_limit f.1 g.1 -end - -instance forget_to_PresheafedSpace_preserves_open_immersion : - PresheafedSpace.is_open_immersion ((LocallyRingedSpace.forget_to_SheafedSpace ⋙ - SheafedSpace.forget_to_PresheafedSpace).map f) := H - -instance forget_to_Top_preserves_pullback_of_left : - preserves_limit (cospan f g) - (LocallyRingedSpace.forget_to_SheafedSpace ⋙ SheafedSpace.forget _) := -begin - change preserves_limit _ - ((LocallyRingedSpace.forget_to_SheafedSpace ⋙ SheafedSpace.forget_to_PresheafedSpace) - ⋙ PresheafedSpace.forget _), - apply_with limits.comp_preserves_limit { instances := ff }, - apply_instance, - apply preserves_limit_of_iso_diagram _ (diagram_iso_cospan.{u} _).symm, - dsimp [SheafedSpace.forget_to_PresheafedSpace], - apply_instance, -end - -instance forget_reflects_pullback_of_left : - reflects_limit (cospan f g) LocallyRingedSpace.forget_to_SheafedSpace := -reflects_limit_of_reflects_isomorphisms _ _ - -instance forget_preserves_pullback_of_right : - preserves_limit (cospan g f) LocallyRingedSpace.forget_to_SheafedSpace := -preserves_pullback_symmetry _ _ _ - -instance forget_to_PresheafedSpace_preserves_pullback_of_right : - preserves_limit (cospan g f) (LocallyRingedSpace.forget_to_SheafedSpace ⋙ - SheafedSpace.forget_to_PresheafedSpace) := -preserves_pullback_symmetry _ _ _ - -instance forget_reflects_pullback_of_right : - reflects_limit (cospan g f) LocallyRingedSpace.forget_to_SheafedSpace := -reflects_limit_of_reflects_isomorphisms _ _ - -instance forget_to_PresheafedSpace_reflects_pullback_of_left : - reflects_limit (cospan f g) - (LocallyRingedSpace.forget_to_SheafedSpace ⋙ SheafedSpace.forget_to_PresheafedSpace) := -reflects_limit_of_reflects_isomorphisms _ _ - -instance forget_to_PresheafedSpace_reflects_pullback_of_right : - reflects_limit (cospan g f) - (LocallyRingedSpace.forget_to_SheafedSpace ⋙ SheafedSpace.forget_to_PresheafedSpace) := -reflects_limit_of_reflects_isomorphisms _ _ - -lemma pullback_snd_is_iso_of_range_subset (H' : set.range g.1.base ⊆ set.range f.1.base) : - is_iso (pullback.snd : pullback f g ⟶ _) := -begin - apply_with (reflects_isomorphisms.reflects LocallyRingedSpace.forget_to_SheafedSpace) - { instances := ff }, - apply_with (reflects_isomorphisms.reflects SheafedSpace.forget_to_PresheafedSpace) - { instances := ff }, - erw ← preserves_pullback.iso_hom_snd - (LocallyRingedSpace.forget_to_SheafedSpace ⋙ SheafedSpace.forget_to_PresheafedSpace) f g, - haveI := PresheafedSpace.is_open_immersion.pullback_snd_is_iso_of_range_subset _ _ H', - apply_instance, - apply_instance -end - -/-- -The universal property of open immersions: -For an open immersion `f : X ⟶ Z`, given any morphism of schemes `g : Y ⟶ Z` whose topological -image is contained in the image of `f`, we can lift this morphism to a unique `Y ⟶ X` that -commutes with these maps. --/ -def lift (H' : set.range g.1.base ⊆ set.range f.1.base) : Y ⟶ X := -begin - haveI := pullback_snd_is_iso_of_range_subset f g H', - exact inv (pullback.snd : pullback f g ⟶ _) ≫ pullback.fst, -end - -@[simp, reassoc] lemma lift_fac (H' : set.range g.1.base ⊆ set.range f.1.base) : - lift f g H' ≫ f = g := -by { erw category.assoc, rw is_iso.inv_comp_eq, exact pullback.condition } - -lemma lift_uniq (H' : set.range g.1.base ⊆ set.range f.1.base) (l : Y ⟶ X) - (hl : l ≫ f = g) : l = lift f g H' := -by rw [← cancel_mono f, hl, lift_fac] - -lemma lift_range (H' : set.range g.1.base ⊆ set.range f.1.base) : - set.range (lift f g H').1.base = f.1.base ⁻¹' (set.range g.1.base) := -begin - haveI := pullback_snd_is_iso_of_range_subset f g H', - dsimp only [lift], - have : _ = (pullback.fst : pullback f g ⟶ _).val.base := preserves_pullback.iso_hom_fst - (LocallyRingedSpace.forget_to_SheafedSpace ⋙ SheafedSpace.forget _) f g, - rw [LocallyRingedSpace.comp_val, SheafedSpace.comp_base, ← this, ← category.assoc, coe_comp], - rw [set.range_comp, set.range_iff_surjective.mpr, set.image_univ, Top.pullback_fst_range], - ext, - split, - { rintros ⟨y, eq⟩, exact ⟨y, eq.symm⟩ }, - { rintros ⟨y, eq⟩, exact ⟨y, eq.symm⟩ }, - { rw ← Top.epi_iff_surjective, - rw (show (inv (pullback.snd : pullback f g ⟶ _)).val.base = _, from - (LocallyRingedSpace.forget_to_SheafedSpace ⋙ SheafedSpace.forget _).map_inv _), - apply_instance } -end - -end pullback - -/-- An open immersion is isomorphic to the induced open subscheme on its image. -/ -def iso_restrict {X Y : LocallyRingedSpace} {f : X ⟶ Y} - (H : LocallyRingedSpace.is_open_immersion f) : X ≅ Y.restrict H.base_open := -begin - apply LocallyRingedSpace.iso_of_SheafedSpace_iso, - refine SheafedSpace.forget_to_PresheafedSpace.preimage_iso _, - exact H.iso_restrict -end - /-- To show that a locally ringed space is a scheme, it suffices to show that it has a jointly surjective family of open immersions from affine schemes. -/ protected def Scheme (X : LocallyRingedSpace) diff --git a/src/algebraic_geometry/open_immersion/basic.lean b/src/algebraic_geometry/open_immersion/basic.lean new file mode 100644 index 0000000000000..c4c7b15bb57de --- /dev/null +++ b/src/algebraic_geometry/open_immersion/basic.lean @@ -0,0 +1,1081 @@ +/- +Copyright (c) 2021 Andrew Yang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Andrew Yang +-/ +import topology.category.Top.limits.pullbacks +import algebraic_geometry.locally_ringed_space + +/-! +# Open immersions of structured spaces + +We say that a morphism of presheafed spaces `f : X ⟶ Y` is an open immersion if +the underlying map of spaces is an open embedding `f : X ⟶ U ⊆ Y`, +and the sheaf map `Y(V) ⟶ f _* X(V)` is an iso for each `V ⊆ U`. + +Abbreviations are also provided for `SheafedSpace`, `LocallyRingedSpace` and `Scheme`. + +## Main definitions + +* `algebraic_geometry.PresheafedSpace.is_open_immersion`: the `Prop`-valued typeclass asserting + that a PresheafedSpace hom `f` is an open_immersion. +* `algebraic_geometry.is_open_immersion`: the `Prop`-valued typeclass asserting + that a Scheme morphism `f` is an open_immersion. +* `algebraic_geometry.PresheafedSpace.is_open_immersion.iso_restrict`: The source of an + open immersion is isomorphic to the restriction of the target onto the image. +* `algebraic_geometry.PresheafedSpace.is_open_immersion.lift`: Any morphism whose range is + contained in an open immersion factors though the open immersion. +* `algebraic_geometry.PresheafedSpace.is_open_immersion.to_SheafedSpace`: If `f : X ⟶ Y` is an + open immersion of presheafed spaces, and `Y` is a sheafed space, then `X` is also a sheafed + space. The morphism as morphisms of sheafed spaces is given by `to_SheafedSpace_hom`. +* `algebraic_geometry.PresheafedSpace.is_open_immersion.to_LocallyRingedSpace`: If `f : X ⟶ Y` is + an open immersion of presheafed spaces, and `Y` is a locally ringed space, then `X` is also a + locally ringed space. The morphism as morphisms of locally ringed spaces is given by + `to_LocallyRingedSpace_hom`. + +## Main results + +* `algebraic_geometry.PresheafedSpace.is_open_immersion.comp`: The composition of two open + immersions is an open immersion. +* `algebraic_geometry.PresheafedSpace.is_open_immersion.of_iso`: An iso is an open immersion. +* `algebraic_geometry.PresheafedSpace.is_open_immersion.to_iso`: + A surjective open immersion is an isomorphism. +* `algebraic_geometry.PresheafedSpace.is_open_immersion.stalk_iso`: An open immersion induces + an isomorphism on stalks. +* `algebraic_geometry.PresheafedSpace.is_open_immersion.has_pullback_of_left`: If `f` is an open + immersion, then the pullback `(f, g)` exists (and the forgetful functor to `Top` preserves it). +* `algebraic_geometry.PresheafedSpace.is_open_immersion.pullback_snd_of_left`: Open immersions + are stable under pullbacks. +* `algebraic_geometry.SheafedSpace.is_open_immersion.of_stalk_iso` An (topological) open embedding + between two sheafed spaces is an open immersion if all the stalk maps are isomorphisms. + +-/ + +open topological_space category_theory opposite +open category_theory.limits +namespace algebraic_geometry + +universes v v₁ v₂ u + +variables {C : Type u} [category.{v} C] + +/-- +An open immersion of PresheafedSpaces is an open embedding `f : X ⟶ U ⊆ Y` of the underlying +spaces, such that the sheaf map `Y(V) ⟶ f _* X(V)` is an iso for each `V ⊆ U`. +-/ +class PresheafedSpace.is_open_immersion {X Y : PresheafedSpace.{v} C} (f : X ⟶ Y) : Prop := +(base_open : open_embedding f.base) +(c_iso : ∀ U : opens X, is_iso (f.c.app (op (base_open.is_open_map.functor.obj U)))) + +/-- +A morphism of SheafedSpaces is an open immersion if it is an open immersion as a morphism +of PresheafedSpaces +-/ +abbreviation SheafedSpace.is_open_immersion {X Y : SheafedSpace.{v} C} (f : X ⟶ Y) : Prop := +PresheafedSpace.is_open_immersion f + +/-- +A morphism of LocallyRingedSpaces is an open immersion if it is an open immersion as a morphism +of SheafedSpaces +-/ +abbreviation LocallyRingedSpace.is_open_immersion {X Y : LocallyRingedSpace} (f : X ⟶ Y) : Prop := +SheafedSpace.is_open_immersion f.1 + +namespace PresheafedSpace.is_open_immersion + +open PresheafedSpace + +local notation `is_open_immersion` := PresheafedSpace.is_open_immersion + +attribute [instance] is_open_immersion.c_iso + +section + +variables {X Y : PresheafedSpace.{v} C} {f : X ⟶ Y} (H : is_open_immersion f) + +/-- The functor `opens X ⥤ opens Y` associated with an open immersion `f : X ⟶ Y`. -/ +abbreviation open_functor := H.base_open.is_open_map.functor + +/-- An open immersion `f : X ⟶ Y` induces an isomorphism `X ≅ Y|_{f(X)}`. -/ +@[simps hom_c_app] noncomputable +def iso_restrict : X ≅ Y.restrict H.base_open := +PresheafedSpace.iso_of_components (iso.refl _) +begin + symmetry, + fapply nat_iso.of_components, + intro U, + refine as_iso (f.c.app (op (H.open_functor.obj (unop U)))) ≪≫ X.presheaf.map_iso (eq_to_iso _), + { induction U using opposite.rec, + cases U, + dsimp only [is_open_map.functor, functor.op, opens.map], + congr' 2, + erw set.preimage_image_eq _ H.base_open.inj, + refl }, + { intros U V i, + simp only [category_theory.eq_to_iso.hom, Top.presheaf.pushforward_obj_map, category.assoc, + functor.op_map, iso.trans_hom, as_iso_hom, functor.map_iso_hom, ←X.presheaf.map_comp], + erw [f.c.naturality_assoc, ←X.presheaf.map_comp], + congr } +end + +@[simp] lemma iso_restrict_hom_of_restrict : H.iso_restrict.hom ≫ Y.of_restrict _ = f := +begin + ext, + { simp only [comp_c_app, iso_restrict_hom_c_app, nat_trans.comp_app, + eq_to_hom_refl, of_restrict_c_app, category.assoc, whisker_right_id'], + erw [category.comp_id, f.c.naturality_assoc, ←X.presheaf.map_comp], + transitivity f.c.app x ≫ X.presheaf.map (𝟙 _), + { congr }, + { erw [X.presheaf.map_id, category.comp_id] } }, + { refl, } +end + +@[simp] lemma iso_restrict_inv_of_restrict : H.iso_restrict.inv ≫ f = Y.of_restrict _ := +by { rw [iso.inv_comp_eq, iso_restrict_hom_of_restrict] } + +instance mono [H : is_open_immersion f] : mono f := +by { rw ← H.iso_restrict_hom_of_restrict, apply mono_comp } + +/-- The composition of two open immersions is an open immersion. -/ +instance comp {Z : PresheafedSpace C} (f : X ⟶ Y) [hf : is_open_immersion f] (g : Y ⟶ Z) + [hg : is_open_immersion g] : + is_open_immersion (f ≫ g) := +{ base_open := hg.base_open.comp hf.base_open, + c_iso := λ U, + begin + generalize_proofs h, + dsimp only [algebraic_geometry.PresheafedSpace.comp_c_app, unop_op, functor.op, comp_base, + Top.presheaf.pushforward_obj_obj, opens.map_comp_obj], + apply_with is_iso.comp_is_iso { instances := ff }, + swap, + { have : (opens.map g.base).obj (h.functor.obj U) = hf.open_functor.obj U, + { ext1, + dsimp only [opens.map_coe, is_open_map.functor_obj_coe, comp_base], + rw [coe_comp, ← set.image_image, set.preimage_image_eq _ hg.base_open.inj] }, + rw this, + apply_instance }, + { have : h.functor.obj U = hg.open_functor.obj (hf.open_functor.obj U), + { ext1, + dsimp only [is_open_map.functor_obj_coe], + rw [comp_base, coe_comp, ←set.image_image] }, + rw this, + apply_instance } + end } + +/-- For an open immersion `f : X ⟶ Y` and an open set `U ⊆ X`, we have the map `X(U) ⟶ Y(U)`. -/ +noncomputable +def inv_app (U : opens X) : X.presheaf.obj (op U) ⟶ Y.presheaf.obj (op (H.open_functor.obj U)) := +X.presheaf.map (eq_to_hom (by simp [opens.map, set.preimage_image_eq _ H.base_open.inj])) ≫ + inv (f.c.app (op (H.open_functor.obj U))) + +@[simp, reassoc] lemma inv_naturality {U V : (opens X)ᵒᵖ} (i : U ⟶ V) : + X.presheaf.map i ≫ H.inv_app (unop V) = H.inv_app (unop U) ≫ + Y.presheaf.map (H.open_functor.op.map i) := +begin + simp only [inv_app, ←category.assoc], + rw [is_iso.comp_inv_eq], + simp only [category.assoc, f.c.naturality, is_iso.inv_hom_id_assoc, ← X.presheaf.map_comp], + erw ← X.presheaf.map_comp, + congr +end + +instance (U : opens X) : is_iso (H.inv_app U) := by { delta inv_app, apply_instance } + +lemma inv_inv_app (U : opens X) : + inv (H.inv_app U) = f.c.app (op (H.open_functor.obj U)) ≫ + X.presheaf.map (eq_to_hom (by simp [opens.map, set.preimage_image_eq _ H.base_open.inj])) := +begin + rw ← cancel_epi (H.inv_app U), + rw is_iso.hom_inv_id, + delta inv_app, + simp [← functor.map_comp] +end + +@[simp, reassoc, elementwise] lemma inv_app_app (U : opens X) : + H.inv_app U ≫ f.c.app (op (H.open_functor.obj U)) = + X.presheaf.map (eq_to_hom (by simp [opens.map, set.preimage_image_eq _ H.base_open.inj])) := +by rw [inv_app, category.assoc, is_iso.inv_hom_id, category.comp_id] + +@[simp, reassoc] lemma app_inv_app (U : opens Y) : + f.c.app (op U) ≫ H.inv_app ((opens.map f.base).obj U) = + Y.presheaf.map ((hom_of_le (by exact set.image_preimage_subset f.base U)).op : + op U ⟶ op (H.open_functor.obj ((opens.map f.base).obj U))) := +by { erw ← category.assoc, rw [is_iso.comp_inv_eq, f.c.naturality], congr } + +/-- A variant of `app_inv_app` that gives an `eq_to_hom` instead of `hom_of_le`. -/ +@[reassoc] lemma app_inv_app' (U : opens Y) (hU : (U : set Y) ⊆ set.range f.base) : + f.c.app (op U) ≫ H.inv_app ((opens.map f.base).obj U) = + Y.presheaf.map (eq_to_hom (by + { apply le_antisymm, + { exact set.image_preimage_subset f.base U.1 }, + { rw [← set_like.coe_subset_coe], + refine has_le.le.trans_eq _ (@set.image_preimage_eq_inter_range _ _ f.base U.1).symm, + exact set.subset_inter_iff.mpr ⟨λ _ h, h, hU⟩ } })).op := +by { erw ← category.assoc, rw [is_iso.comp_inv_eq, f.c.naturality], congr } + +/-- An isomorphism is an open immersion. -/ +instance of_iso {X Y : PresheafedSpace.{v} C} (H : X ≅ Y) : is_open_immersion H.hom := +{ base_open := (Top.homeo_of_iso ((forget C).map_iso H)).open_embedding, + c_iso := λ _, infer_instance } + +@[priority 100] +instance of_is_iso {X Y : PresheafedSpace.{v} C} (f : X ⟶ Y) [is_iso f] : is_open_immersion f := +algebraic_geometry.PresheafedSpace.is_open_immersion.of_iso (as_iso f) + +instance of_restrict {X : Top} (Y : PresheafedSpace C) {f : X ⟶ Y.carrier} + (hf : open_embedding f) : is_open_immersion (Y.of_restrict hf) := +{ base_open := hf, + c_iso := λ U, + begin + dsimp, + have : (opens.map f).obj (hf.is_open_map.functor.obj U) = U, + { ext1, + exact set.preimage_image_eq _ hf.inj }, + convert (show is_iso (Y.presheaf.map (𝟙 _)), from infer_instance), + { apply subsingleton.helim, + rw this }, + { rw Y.presheaf.map_id, + apply_instance } + end } + +@[elementwise, simp] +lemma of_restrict_inv_app {C : Type*} [category C] (X : PresheafedSpace C) {Y : Top} + {f : Y ⟶ Top.of X.carrier} + (h : open_embedding f) (U : opens (X.restrict h).carrier) : + (PresheafedSpace.is_open_immersion.of_restrict X h).inv_app U = 𝟙 _ := +begin + delta PresheafedSpace.is_open_immersion.inv_app, + rw [is_iso.comp_inv_eq, category.id_comp], + change X.presheaf.map _ = X.presheaf.map _, + congr +end + +/-- An open immersion is an iso if the underlying continuous map is epi. -/ +lemma to_iso (f : X ⟶ Y) [h : is_open_immersion f] [h' : epi f.base] : is_iso f := +begin + apply_with is_iso_of_components { instances := ff }, + { let : X ≃ₜ Y := (homeomorph.of_embedding _ h.base_open.to_embedding).trans + { to_fun := subtype.val, inv_fun := λ x, ⟨x, + by { rw set.range_iff_surjective.mpr ((Top.epi_iff_surjective _).mp h'), trivial }⟩, + left_inv := λ ⟨_,_⟩, rfl, right_inv := λ _, rfl }, + convert is_iso.of_iso (Top.iso_of_homeo this), + { ext, refl } }, + { apply_with nat_iso.is_iso_of_is_iso_app { instances := ff }, + intro U, + have : U = op (h.open_functor.obj ((opens.map f.base).obj (unop U))), + { induction U using opposite.rec, + cases U, + dsimp only [functor.op, opens.map], + congr, + exact (set.image_preimage_eq _ ((Top.epi_iff_surjective _).mp h')).symm }, + convert @@is_open_immersion.c_iso _ h ((opens.map f.base).obj (unop U)) } +end + +instance stalk_iso [has_colimits C] [H : is_open_immersion f] (x : X) : is_iso (stalk_map f x) := +begin + rw ← H.iso_restrict_hom_of_restrict, + rw PresheafedSpace.stalk_map.comp, + apply_instance +end + +end + +section pullback + +noncomputable theory + +variables {X Y Z : PresheafedSpace.{v} C} (f : X ⟶ Z) [hf : is_open_immersion f] (g : Y ⟶ Z) + +include hf + +/-- + (Implementation.) The projection map when constructing the pullback along an open immersion. +-/ +def pullback_cone_of_left_fst : + Y.restrict (Top.snd_open_embedding_of_left_open_embedding hf.base_open g.base) ⟶ X := +{ base := pullback.fst, + c := + { app := λ U, hf.inv_app (unop U) ≫ + g.c.app (op (hf.base_open.is_open_map.functor.obj (unop U))) ≫ + Y.presheaf.map (eq_to_hom + (begin + simp only [is_open_map.functor, subtype.mk_eq_mk, unop_op, op_inj_iff, opens.map, + subtype.coe_mk, functor.op_obj, subtype.val_eq_coe], + apply has_le.le.antisymm, + { rintros _ ⟨_, h₁, h₂⟩, + use (Top.pullback_iso_prod_subtype _ _).inv ⟨⟨_, _⟩, h₂⟩, + simpa using h₁ }, + { rintros _ ⟨x, h₁, rfl⟩, + exact ⟨_, h₁, concrete_category.congr_hom pullback.condition x⟩ } + end)), + naturality' := + begin + intros U V i, + induction U using opposite.rec, + induction V using opposite.rec, + simp only [quiver.hom.unop_op, Top.presheaf.pushforward_obj_map, category.assoc, + nat_trans.naturality_assoc, functor.op_map, inv_naturality_assoc, ← Y.presheaf.map_comp], + erw ← Y.presheaf.map_comp, + congr + end } } + +lemma pullback_cone_of_left_condition : + pullback_cone_of_left_fst f g ≫ f = Y.of_restrict _ ≫ g := +begin + ext U, + { induction U using opposite.rec, + dsimp only [comp_c_app, nat_trans.comp_app, unop_op, + whisker_right_app, pullback_cone_of_left_fst], + simp only [quiver.hom.unop_op, Top.presheaf.pushforward_obj_map, app_inv_app_assoc, + eq_to_hom_app, eq_to_hom_unop, category.assoc, nat_trans.naturality_assoc, functor.op_map], + erw [← Y.presheaf.map_comp, ← Y.presheaf.map_comp], + congr }, + { simpa using pullback.condition } +end + +/-- +We construct the pullback along an open immersion via restricting along the pullback of the +maps of underlying spaces (which is also an open embedding). +-/ +def pullback_cone_of_left : pullback_cone f g := +pullback_cone.mk (pullback_cone_of_left_fst f g) (Y.of_restrict _) + (pullback_cone_of_left_condition f g) + +variable (s : pullback_cone f g) + +/-- + (Implementation.) Any cone over `cospan f g` indeed factors through the constructed cone. +-/ +def pullback_cone_of_left_lift : s.X ⟶ (pullback_cone_of_left f g).X := +{ base := pullback.lift s.fst.base s.snd.base + (congr_arg (λ x, PresheafedSpace.hom.base x) s.condition), + c := + { app := λ U, s.snd.c.app _ ≫ s.X.presheaf.map (eq_to_hom (begin + dsimp only [opens.map, is_open_map.functor, functor.op], + congr' 2, + let s' : pullback_cone f.base g.base := pullback_cone.mk s.fst.base s.snd.base _, + have : _ = s.snd.base := limit.lift_π s' walking_cospan.right, + conv_lhs { erw ← this, rw coe_comp, erw ← set.preimage_preimage }, + erw set.preimage_image_eq _ + (Top.snd_open_embedding_of_left_open_embedding hf.base_open g.base).inj, + end)), + naturality' := λ U V i, + begin + erw s.snd.c.naturality_assoc, + rw category.assoc, + erw [← s.X.presheaf.map_comp, ← s.X.presheaf.map_comp], + congr + end } } + +-- this lemma is not a `simp` lemma, because it is an implementation detail +lemma pullback_cone_of_left_lift_fst : + pullback_cone_of_left_lift f g s ≫ (pullback_cone_of_left f g).fst = s.fst := +begin + ext x, + { induction x using opposite.rec, + change ((_ ≫ _) ≫ _ ≫ _) ≫ _ = _, + simp_rw [category.assoc], + erw ← s.X.presheaf.map_comp, + erw s.snd.c.naturality_assoc, + have := congr_app s.condition (op (hf.open_functor.obj x)), + dsimp only [comp_c_app, unop_op] at this, + rw ← is_iso.comp_inv_eq at this, + reassoc! this, + erw [← this, hf.inv_app_app_assoc, s.fst.c.naturality_assoc], + simpa [eq_to_hom_map], }, + { change pullback.lift _ _ _ ≫ pullback.fst = _, + simp } +end + +-- this lemma is not a `simp` lemma, because it is an implementation detail +lemma pullback_cone_of_left_lift_snd : + pullback_cone_of_left_lift f g s ≫ (pullback_cone_of_left f g).snd = s.snd := +begin + ext x, + { change (_ ≫ _ ≫ _) ≫ _ = _, + simp_rw category.assoc, + erw s.snd.c.naturality_assoc, + erw [← s.X.presheaf.map_comp, ← s.X.presheaf.map_comp], + transitivity s.snd.c.app x ≫ s.X.presheaf.map (𝟙 _), + { congr }, + { rw s.X.presheaf.map_id, erw category.comp_id } }, + { change pullback.lift _ _ _ ≫ pullback.snd = _, + simp } +end + +instance pullback_cone_snd_is_open_immersion : + is_open_immersion (pullback_cone_of_left f g).snd := +begin + erw category_theory.limits.pullback_cone.mk_snd, + apply_instance +end + +/-- The constructed pullback cone is indeed the pullback. -/ +def pullback_cone_of_left_is_limit : + is_limit (pullback_cone_of_left f g) := +begin + apply pullback_cone.is_limit_aux', + intro s, + use pullback_cone_of_left_lift f g s, + use pullback_cone_of_left_lift_fst f g s, + use pullback_cone_of_left_lift_snd f g s, + intros m h₁ h₂, + rw ← cancel_mono (pullback_cone_of_left f g).snd, + exact (h₂.trans (pullback_cone_of_left_lift_snd f g s).symm) +end + +instance has_pullback_of_left : + has_pullback f g := +⟨⟨⟨_, pullback_cone_of_left_is_limit f g⟩⟩⟩ + +instance has_pullback_of_right : + has_pullback g f := has_pullback_symmetry f g + +/-- Open immersions are stable under base-change. -/ +instance pullback_snd_of_left : + is_open_immersion (pullback.snd : pullback f g ⟶ _) := +begin + delta pullback.snd, + rw ← limit.iso_limit_cone_hom_π ⟨_, pullback_cone_of_left_is_limit f g⟩ walking_cospan.right, + apply_instance +end + +/-- Open immersions are stable under base-change. -/ +instance pullback_fst_of_right : + is_open_immersion (pullback.fst : pullback g f ⟶ _) := +begin + rw ← pullback_symmetry_hom_comp_snd, + apply_instance +end + +instance pullback_to_base_is_open_immersion [is_open_immersion g] : + is_open_immersion (limit.π (cospan f g) walking_cospan.one) := +begin + rw [←limit.w (cospan f g) walking_cospan.hom.inl, cospan_map_inl], + apply_instance +end + +instance forget_preserves_limits_of_left : preserves_limit (cospan f g) (forget C) := +preserves_limit_of_preserves_limit_cone (pullback_cone_of_left_is_limit f g) +begin + apply (is_limit.postcompose_hom_equiv (diagram_iso_cospan.{v} _) _).to_fun, + refine (is_limit.equiv_iso_limit _).to_fun (limit.is_limit (cospan f.base g.base)), + fapply cones.ext, + exact (iso.refl _), + change ∀ j, _ = 𝟙 _ ≫ _ ≫ _, + simp_rw category.id_comp, + rintros (_|_|_); symmetry, + { erw category.comp_id, + exact limit.w (cospan f.base g.base) walking_cospan.hom.inl }, + { exact category.comp_id _ }, + { exact category.comp_id _ }, +end + +instance forget_preserves_limits_of_right : preserves_limit (cospan g f) (forget C) := +preserves_pullback_symmetry (forget C) f g + +lemma pullback_snd_is_iso_of_range_subset (H : set.range g.base ⊆ set.range f.base) : + is_iso (pullback.snd : pullback f g ⟶ _) := +begin + haveI := Top.snd_iso_of_left_embedding_range_subset hf.base_open.to_embedding g.base H, + haveI : is_iso (pullback.snd : pullback f g ⟶ _).base, + { delta pullback.snd, + rw ← limit.iso_limit_cone_hom_π ⟨_, pullback_cone_of_left_is_limit f g⟩ walking_cospan.right, + change is_iso (_ ≫ pullback.snd), + apply_instance }, + apply to_iso +end + +/-- +The universal property of open immersions: +For an open immersion `f : X ⟶ Z`, given any morphism of schemes `g : Y ⟶ Z` whose topological +image is contained in the image of `f`, we can lift this morphism to a unique `Y ⟶ X` that +commutes with these maps. +-/ +def lift (H : set.range g.base ⊆ set.range f.base) : Y ⟶ X := +begin + haveI := pullback_snd_is_iso_of_range_subset f g H, + exact inv (pullback.snd : pullback f g ⟶ _) ≫ pullback.fst, +end + +@[simp, reassoc] lemma lift_fac (H : set.range g.base ⊆ set.range f.base) : + lift f g H ≫ f = g := +by { erw category.assoc, rw is_iso.inv_comp_eq, exact pullback.condition } + +lemma lift_uniq (H : set.range g.base ⊆ set.range f.base) (l : Y ⟶ X) + (hl : l ≫ f = g) : l = lift f g H := +by rw [← cancel_mono f, hl, lift_fac] + +/-- Two open immersions with equal range is isomorphic. -/ +@[simps] def iso_of_range_eq [is_open_immersion g] (e : set.range f.base = set.range g.base) : + X ≅ Y := +{ hom := lift g f (le_of_eq e), + inv := lift f g (le_of_eq e.symm), + hom_inv_id' := by { rw ← cancel_mono f, simp }, + inv_hom_id' := by { rw ← cancel_mono g, simp } } + +end pullback + +open category_theory.limits.walking_cospan + +section to_SheafedSpace + +variables {X : PresheafedSpace.{v} C} (Y : SheafedSpace C) +variables (f : X ⟶ Y.to_PresheafedSpace) [H : is_open_immersion f] + +include H + +/-- If `X ⟶ Y` is an open immersion, and `Y` is a SheafedSpace, then so is `X`. -/ +def to_SheafedSpace : SheafedSpace C := +{ is_sheaf := + begin + apply Top.presheaf.is_sheaf_of_iso (sheaf_iso_of_iso H.iso_restrict.symm).symm, + apply Top.sheaf.pushforward_sheaf_of_sheaf, + exact (Y.restrict H.base_open).is_sheaf + end, + to_PresheafedSpace := X } + +@[simp] lemma to_SheafedSpace_to_PresheafedSpace : (to_SheafedSpace Y f).to_PresheafedSpace = X := +rfl + +/-- +If `X ⟶ Y` is an open immersion of PresheafedSpaces, and `Y` is a SheafedSpace, we can +upgrade it into a morphism of SheafedSpaces. +-/ +def to_SheafedSpace_hom : to_SheafedSpace Y f ⟶ Y := f + +@[simp] lemma to_SheafedSpace_hom_base : (to_SheafedSpace_hom Y f).base = f.base := rfl + +@[simp] lemma to_SheafedSpace_hom_c : (to_SheafedSpace_hom Y f).c = f.c := rfl + +instance to_SheafedSpace_is_open_immersion : + SheafedSpace.is_open_immersion (to_SheafedSpace_hom Y f) := H + +omit H + +@[simp] lemma SheafedSpace_to_SheafedSpace {X Y : SheafedSpace.{v} C} (f : X ⟶ Y) + [is_open_immersion f] : to_SheafedSpace Y f = X := by unfreezingI { cases X, refl } + +end to_SheafedSpace + +section to_LocallyRingedSpace + +variables {X : PresheafedSpace.{u} CommRing.{u}} (Y : LocallyRingedSpace.{u}) +variables (f : X ⟶ Y.to_PresheafedSpace) [H : is_open_immersion f] + +include H + +/-- If `X ⟶ Y` is an open immersion, and `Y` is a LocallyRingedSpace, then so is `X`. -/ +def to_LocallyRingedSpace : LocallyRingedSpace := +{ to_SheafedSpace := to_SheafedSpace Y.to_SheafedSpace f, + local_ring := λ x, begin + haveI : local_ring (Y.to_SheafedSpace.to_PresheafedSpace.stalk (f.base x)) := Y.local_ring _, + exact (as_iso (stalk_map f x)).CommRing_iso_to_ring_equiv.local_ring + end } + +@[simp] lemma to_LocallyRingedSpace_to_SheafedSpace : + (to_LocallyRingedSpace Y f).to_SheafedSpace = (to_SheafedSpace Y.1 f) := rfl + +/-- +If `X ⟶ Y` is an open immersion of PresheafedSpaces, and `Y` is a LocallyRingedSpace, we can +upgrade it into a morphism of LocallyRingedSpace. +-/ +def to_LocallyRingedSpace_hom : to_LocallyRingedSpace Y f ⟶ Y := ⟨f, λ x, infer_instance⟩ + +@[simp] lemma to_LocallyRingedSpace_hom_val : + (to_LocallyRingedSpace_hom Y f).val = f := rfl + +instance to_LocallyRingedSpace_is_open_immersion : + LocallyRingedSpace.is_open_immersion (to_LocallyRingedSpace_hom Y f) := H + +omit H + +@[simp] lemma LocallyRingedSpace_to_LocallyRingedSpace {X Y : LocallyRingedSpace} (f : X ⟶ Y) + [LocallyRingedSpace.is_open_immersion f] : + to_LocallyRingedSpace Y f.1 = X := +by unfreezingI { cases X, delta to_LocallyRingedSpace, simp } + +end to_LocallyRingedSpace + +lemma is_iso_of_subset {X Y : PresheafedSpace.{v} C} (f : X ⟶ Y) + [H : PresheafedSpace.is_open_immersion f] (U : opens Y.carrier) + (hU : (U : set Y.carrier) ⊆ set.range f.base) : is_iso (f.c.app $ op U) := +begin + have : U = H.base_open.is_open_map.functor.obj ((opens.map f.base).obj U), + { ext1, + exact (set.inter_eq_left_iff_subset.mpr hU).symm.trans set.image_preimage_eq_inter_range.symm }, + convert PresheafedSpace.is_open_immersion.c_iso ((opens.map f.base).obj U), +end + +end PresheafedSpace.is_open_immersion + +namespace SheafedSpace.is_open_immersion + +@[priority 100] +instance of_is_iso {X Y : SheafedSpace.{v} C} (f : X ⟶ Y) [is_iso f] : + SheafedSpace.is_open_immersion f := +@@PresheafedSpace.is_open_immersion.of_is_iso _ f +(SheafedSpace.forget_to_PresheafedSpace.map_is_iso _) + +instance comp {X Y Z : SheafedSpace C} (f : X ⟶ Y) (g : Y ⟶ Z) + [SheafedSpace.is_open_immersion f] [SheafedSpace.is_open_immersion g] : + SheafedSpace.is_open_immersion (f ≫ g) := PresheafedSpace.is_open_immersion.comp f g + +section pullback + +variables {X Y Z : SheafedSpace C} (f : X ⟶ Z) (g : Y ⟶ Z) +variable [H : SheafedSpace.is_open_immersion f] + +include H + +local notation `forget` := SheafedSpace.forget_to_PresheafedSpace +open category_theory.limits.walking_cospan + +instance : mono f := +forget .mono_of_mono_map (show @mono (PresheafedSpace C) _ _ _ f, by apply_instance) + +instance forget_map_is_open_immersion : + PresheafedSpace.is_open_immersion (forget .map f) := ⟨H.base_open, H.c_iso⟩ + +instance has_limit_cospan_forget_of_left : has_limit (cospan f g ⋙ forget) := +begin + apply has_limit_of_iso (diagram_iso_cospan.{v} _).symm, + change has_limit (cospan (forget .map f) (forget .map g)), + apply_instance +end + +instance has_limit_cospan_forget_of_left' : has_limit (cospan ((cospan f g ⋙ forget).map hom.inl) + ((cospan f g ⋙ forget).map hom.inr)) := +show has_limit (cospan (forget .map f) (forget .map g)), from infer_instance + +instance has_limit_cospan_forget_of_right : has_limit (cospan g f ⋙ forget) := +begin + apply has_limit_of_iso (diagram_iso_cospan.{v} _).symm, + change has_limit (cospan (forget .map g) (forget .map f)), + apply_instance +end + +instance has_limit_cospan_forget_of_right' : has_limit (cospan ((cospan g f ⋙ forget).map hom.inl) + ((cospan g f ⋙ forget).map hom.inr)) := +show has_limit (cospan (forget .map g) (forget .map f)), from infer_instance + + +instance forget_creates_pullback_of_left : creates_limit (cospan f g) forget := +creates_limit_of_fully_faithful_of_iso + (PresheafedSpace.is_open_immersion.to_SheafedSpace Y + (@pullback.snd (PresheafedSpace C) _ _ _ _ f g _)) + (eq_to_iso (show pullback _ _ = pullback _ _, by congr) + ≪≫ has_limit.iso_of_nat_iso (diagram_iso_cospan _).symm) + +instance forget_creates_pullback_of_right : creates_limit (cospan g f) forget := +creates_limit_of_fully_faithful_of_iso + (PresheafedSpace.is_open_immersion.to_SheafedSpace Y + (@pullback.fst (PresheafedSpace C) _ _ _ _ g f _)) + (eq_to_iso (show pullback _ _ = pullback _ _, by congr) + ≪≫ has_limit.iso_of_nat_iso (diagram_iso_cospan _).symm) + +instance SheafedSpace_forget_preserves_of_left : + preserves_limit (cospan f g) (SheafedSpace.forget C) := +@@limits.comp_preserves_limit _ _ _ _ forget (PresheafedSpace.forget C) _ +begin + apply_with (preserves_limit_of_iso_diagram _ (diagram_iso_cospan.{v} _).symm) { instances := tt }, + dsimp, + apply_instance +end + +instance SheafedSpace_forget_preserves_of_right : + preserves_limit (cospan g f) (SheafedSpace.forget C) := +preserves_pullback_symmetry _ _ _ + +instance SheafedSpace_has_pullback_of_left : has_pullback f g := + has_limit_of_created (cospan f g) forget + +instance SheafedSpace_has_pullback_of_right : has_pullback g f := + has_limit_of_created (cospan g f) forget + +/-- Open immersions are stable under base-change. -/ +instance SheafedSpace_pullback_snd_of_left : + SheafedSpace.is_open_immersion (pullback.snd : pullback f g ⟶ _) := +begin + delta pullback.snd, + have : _ = limit.π (cospan f g) right := preserves_limits_iso_hom_π + forget (cospan f g) right, + rw ← this, + have := has_limit.iso_of_nat_iso_hom_π + (diagram_iso_cospan.{v} (cospan f g ⋙ forget)) + right, + erw category.comp_id at this, + rw ← this, + dsimp, + apply_instance +end + +instance SheafedSpace_pullback_fst_of_right : + SheafedSpace.is_open_immersion (pullback.fst : pullback g f ⟶ _) := +begin + delta pullback.fst, + have : _ = limit.π (cospan g f) left := preserves_limits_iso_hom_π + forget (cospan g f) left, + rw ← this, + have := has_limit.iso_of_nat_iso_hom_π + (diagram_iso_cospan.{v} (cospan g f ⋙ forget)) left, + erw category.comp_id at this, + rw ← this, + dsimp, + apply_instance +end + +instance SheafedSpace_pullback_to_base_is_open_immersion [SheafedSpace.is_open_immersion g] : + SheafedSpace.is_open_immersion (limit.π (cospan f g) one : pullback f g ⟶ Z) := +begin + rw [←limit.w (cospan f g) hom.inl, cospan_map_inl], + apply_instance +end + +end pullback + +section of_stalk_iso +variables [has_limits C] [has_colimits C] [concrete_category.{v} C] +variables [reflects_isomorphisms (forget C)] [preserves_limits (forget C)] +variables [preserves_filtered_colimits (forget C)] + +/-- +Suppose `X Y : SheafedSpace C`, where `C` is a concrete category, +whose forgetful functor reflects isomorphisms, preserves limits and filtered colimits. +Then a morphism `X ⟶ Y` that is a topological open embedding +is an open immersion iff every stalk map is an iso. +-/ +lemma of_stalk_iso {X Y : SheafedSpace C} (f : X ⟶ Y) + (hf : open_embedding f.base) [H : ∀ x : X, is_iso (PresheafedSpace.stalk_map f x)] : + SheafedSpace.is_open_immersion f := +{ base_open := hf, + c_iso := λ U, begin + apply_with (Top.presheaf.app_is_iso_of_stalk_functor_map_iso + (show Y.sheaf ⟶ (Top.sheaf.pushforward f.base).obj X.sheaf, from ⟨f.c⟩)) { instances := ff }, + rintros ⟨_, y, hy, rfl⟩, + specialize H y, + delta PresheafedSpace.stalk_map at H, + haveI H' := Top.presheaf.stalk_pushforward.stalk_pushforward_iso_of_open_embedding + C hf X.presheaf y, + have := @@is_iso.comp_is_iso _ H (@@is_iso.inv_is_iso _ H'), + rw [category.assoc, is_iso.hom_inv_id, category.comp_id] at this, + exact this + end } + +end of_stalk_iso + +section prod + +variables [has_limits C] {ι : Type v} (F : discrete ι ⥤ SheafedSpace C) [has_colimit F] + (i : discrete ι) + +lemma sigma_ι_open_embedding : open_embedding (colimit.ι F i).base := +begin + rw ← (show _ = (colimit.ι F i).base, + from ι_preserves_colimits_iso_inv (SheafedSpace.forget C) F i), + have : _ = _ ≫ colimit.ι (discrete.functor ((F ⋙ SheafedSpace.forget C).obj ∘ discrete.mk)) i := + has_colimit.iso_of_nat_iso_ι_hom discrete.nat_iso_functor i, + rw ← iso.eq_comp_inv at this, + rw this, + have : colimit.ι _ _ ≫ _ = _ := + Top.sigma_iso_sigma_hom_ι.{v v} ((F ⋙ SheafedSpace.forget C).obj ∘ discrete.mk) i.as, + rw ← iso.eq_comp_inv at this, + cases i, + rw this, + simp_rw [← category.assoc, Top.open_embedding_iff_comp_is_iso, + Top.open_embedding_iff_is_iso_comp], + dsimp, + exact open_embedding_sigma_mk +end + +lemma image_preimage_is_empty (j : discrete ι) (h : i ≠ j) (U : opens (F.obj i)) : + (opens.map (colimit.ι (F ⋙ SheafedSpace.forget_to_PresheafedSpace) j).base).obj + ((opens.map (preserves_colimit_iso SheafedSpace.forget_to_PresheafedSpace F).inv.base).obj + ((sigma_ι_open_embedding F i).is_open_map.functor.obj U)) = ⊥ := +begin + ext, + apply iff_false_intro, + rintro ⟨y, hy, eq⟩, + replace eq := concrete_category.congr_arg + (preserves_colimit_iso (SheafedSpace.forget C) F ≪≫ + has_colimit.iso_of_nat_iso discrete.nat_iso_functor ≪≫ Top.sigma_iso_sigma.{v} _).hom eq, + simp_rw [category_theory.iso.trans_hom, ← Top.comp_app, ← PresheafedSpace.comp_base] at eq, + rw ι_preserves_colimits_iso_inv at eq, + change ((SheafedSpace.forget C).map (colimit.ι F i) ≫ _) y = + ((SheafedSpace.forget C).map (colimit.ι F j) ≫ _) x at eq, + cases i, cases j, + rw [ι_preserves_colimits_iso_hom_assoc, ι_preserves_colimits_iso_hom_assoc, + has_colimit.iso_of_nat_iso_ι_hom_assoc, has_colimit.iso_of_nat_iso_ι_hom_assoc, + Top.sigma_iso_sigma_hom_ι.{v}, Top.sigma_iso_sigma_hom_ι.{v}] at eq, + exact h (congr_arg discrete.mk (congr_arg sigma.fst eq)), +end + +instance sigma_ι_is_open_immersion [has_strict_terminal_objects C] : + SheafedSpace.is_open_immersion (colimit.ι F i) := +{ base_open := sigma_ι_open_embedding F i, + c_iso := λ U, begin + have e : colimit.ι F i = _ := + (ι_preserves_colimits_iso_inv SheafedSpace.forget_to_PresheafedSpace F i).symm, + have H : open_embedding (colimit.ι (F ⋙ SheafedSpace.forget_to_PresheafedSpace) i ≫ + (preserves_colimit_iso SheafedSpace.forget_to_PresheafedSpace F).inv).base := + e ▸ sigma_ι_open_embedding F i, + suffices : is_iso ((colimit.ι (F ⋙ SheafedSpace.forget_to_PresheafedSpace) i ≫ + (preserves_colimit_iso SheafedSpace.forget_to_PresheafedSpace F).inv).c.app + (op (H.is_open_map.functor.obj U))), + { convert this }, + rw [PresheafedSpace.comp_c_app, + ← PresheafedSpace.colimit_presheaf_obj_iso_componentwise_limit_hom_π], + rsufficesI : is_iso (limit.π (PresheafedSpace.componentwise_diagram + (F ⋙ SheafedSpace.forget_to_PresheafedSpace) + ((opens.map (preserves_colimit_iso SheafedSpace.forget_to_PresheafedSpace F).inv.base).obj + (unop $ op $ H.is_open_map.functor.obj U))) (op i)), + { apply_instance }, + apply limit_π_is_iso_of_is_strict_terminal, + intros j hj, + induction j using opposite.rec, + dsimp, + convert (F.obj j).sheaf.is_terminal_of_empty, + convert image_preimage_is_empty F i j (λ h, hj (congr_arg op h.symm)) U, + exact (congr_arg PresheafedSpace.hom.base e).symm + end } + +end prod + +end SheafedSpace.is_open_immersion + +namespace LocallyRingedSpace.is_open_immersion + +section pullback + +variables {X Y Z : LocallyRingedSpace.{u}} (f : X ⟶ Z) (g : Y ⟶ Z) +variable [H : LocallyRingedSpace.is_open_immersion f] + +@[priority 100] +instance of_is_iso [is_iso g] : + LocallyRingedSpace.is_open_immersion g := +@@PresheafedSpace.is_open_immersion.of_is_iso _ g.1 ⟨⟨(inv g).1, + by { erw ← LocallyRingedSpace.comp_val, rw is_iso.hom_inv_id, + erw ← LocallyRingedSpace.comp_val, rw is_iso.inv_hom_id, split; simpa }⟩⟩ + +include H + +instance comp (g : Z ⟶ Y) [LocallyRingedSpace.is_open_immersion g] : + LocallyRingedSpace.is_open_immersion (f ≫ g) := PresheafedSpace.is_open_immersion.comp f.1 g.1 + +instance mono : mono f := +LocallyRingedSpace.forget_to_SheafedSpace.mono_of_mono_map (show mono f.1, by apply_instance) + +instance : SheafedSpace.is_open_immersion (LocallyRingedSpace.forget_to_SheafedSpace.map f) := H + +/-- An explicit pullback cone over `cospan f g` if `f` is an open immersion. -/ +def pullback_cone_of_left : pullback_cone f g := +begin + refine pullback_cone.mk _ + (Y.of_restrict (Top.snd_open_embedding_of_left_open_embedding H.base_open g.1.base)) _, + { use PresheafedSpace.is_open_immersion.pullback_cone_of_left_fst f.1 g.1, + intro x, + have := PresheafedSpace.stalk_map.congr_hom _ _ + (PresheafedSpace.is_open_immersion.pullback_cone_of_left_condition f.1 g.1) x, + rw [PresheafedSpace.stalk_map.comp, PresheafedSpace.stalk_map.comp] at this, + rw ← is_iso.eq_inv_comp at this, + rw this, + apply_instance }, + { exact LocallyRingedSpace.hom.ext _ _ + (PresheafedSpace.is_open_immersion.pullback_cone_of_left_condition _ _) }, +end + +instance : LocallyRingedSpace.is_open_immersion (pullback_cone_of_left f g).snd := +show PresheafedSpace.is_open_immersion (Y.to_PresheafedSpace.of_restrict _), by apply_instance + +/-- The constructed `pullback_cone_of_left` is indeed limiting. -/ +def pullback_cone_of_left_is_limit : is_limit (pullback_cone_of_left f g) := +pullback_cone.is_limit_aux' _ $ λ s, +begin + use PresheafedSpace.is_open_immersion.pullback_cone_of_left_lift f.1 g.1 + (pullback_cone.mk s.fst.1 s.snd.1 (congr_arg LocallyRingedSpace.hom.val s.condition)), + { intro x, + have := PresheafedSpace.stalk_map.congr_hom _ _ + (PresheafedSpace.is_open_immersion.pullback_cone_of_left_lift_snd f.1 g.1 + (pullback_cone.mk s.fst.1 s.snd.1 (congr_arg LocallyRingedSpace.hom.val s.condition))) x, + change _ = _ ≫ PresheafedSpace.stalk_map s.snd.1 x at this, + rw [PresheafedSpace.stalk_map.comp, ← is_iso.eq_inv_comp] at this, + rw this, + apply_instance }, + split, + { exact LocallyRingedSpace.hom.ext _ _ + (PresheafedSpace.is_open_immersion.pullback_cone_of_left_lift_fst f.1 g.1 _) }, + split, + { exact LocallyRingedSpace.hom.ext _ _ + (PresheafedSpace.is_open_immersion.pullback_cone_of_left_lift_snd f.1 g.1 _) }, + intros m h₁ h₂, + rw ← cancel_mono (pullback_cone_of_left f g).snd, + exact (h₂.trans (LocallyRingedSpace.hom.ext _ _ + (PresheafedSpace.is_open_immersion.pullback_cone_of_left_lift_snd f.1 g.1 + (pullback_cone.mk s.fst.1 s.snd.1 (congr_arg LocallyRingedSpace.hom.val s.condition))).symm)) +end + +instance has_pullback_of_left : + has_pullback f g := +⟨⟨⟨_, pullback_cone_of_left_is_limit f g⟩⟩⟩ + +instance has_pullback_of_right : + has_pullback g f := has_pullback_symmetry f g + +/-- Open immersions are stable under base-change. -/ +instance pullback_snd_of_left : + LocallyRingedSpace.is_open_immersion (pullback.snd : pullback f g ⟶ _) := +begin + delta pullback.snd, + rw ← limit.iso_limit_cone_hom_π ⟨_, pullback_cone_of_left_is_limit f g⟩ walking_cospan.right, + apply_instance +end + +/-- Open immersions are stable under base-change. -/ +instance pullback_fst_of_right : +LocallyRingedSpace.is_open_immersion (pullback.fst : pullback g f ⟶ _) := +begin + rw ← pullback_symmetry_hom_comp_snd, + apply_instance +end + +instance pullback_to_base_is_open_immersion [LocallyRingedSpace.is_open_immersion g] : + LocallyRingedSpace.is_open_immersion (limit.π (cospan f g) walking_cospan.one) := +begin + rw [←limit.w (cospan f g) walking_cospan.hom.inl, cospan_map_inl], + apply_instance +end + +instance forget_preserves_pullback_of_left : + preserves_limit (cospan f g) LocallyRingedSpace.forget_to_SheafedSpace := +preserves_limit_of_preserves_limit_cone (pullback_cone_of_left_is_limit f g) +begin + apply (is_limit_map_cone_pullback_cone_equiv _ _).symm.to_fun, + apply is_limit_of_is_limit_pullback_cone_map SheafedSpace.forget_to_PresheafedSpace, + exact PresheafedSpace.is_open_immersion.pullback_cone_of_left_is_limit f.1 g.1 +end + +instance forget_to_PresheafedSpace_preserves_pullback_of_left : + preserves_limit (cospan f g) + (LocallyRingedSpace.forget_to_SheafedSpace ⋙ SheafedSpace.forget_to_PresheafedSpace) := +preserves_limit_of_preserves_limit_cone (pullback_cone_of_left_is_limit f g) +begin + apply (is_limit_map_cone_pullback_cone_equiv _ _).symm.to_fun, + exact PresheafedSpace.is_open_immersion.pullback_cone_of_left_is_limit f.1 g.1 +end + +instance forget_to_PresheafedSpace_preserves_open_immersion : + PresheafedSpace.is_open_immersion ((LocallyRingedSpace.forget_to_SheafedSpace ⋙ + SheafedSpace.forget_to_PresheafedSpace).map f) := H + +instance forget_to_Top_preserves_pullback_of_left : + preserves_limit (cospan f g) + (LocallyRingedSpace.forget_to_SheafedSpace ⋙ SheafedSpace.forget _) := +begin + change preserves_limit _ + ((LocallyRingedSpace.forget_to_SheafedSpace ⋙ SheafedSpace.forget_to_PresheafedSpace) + ⋙ PresheafedSpace.forget _), + apply_with limits.comp_preserves_limit { instances := ff }, + apply_instance, + apply preserves_limit_of_iso_diagram _ (diagram_iso_cospan.{u} _).symm, + dsimp [SheafedSpace.forget_to_PresheafedSpace], + apply_instance, +end + +instance forget_reflects_pullback_of_left : + reflects_limit (cospan f g) LocallyRingedSpace.forget_to_SheafedSpace := +reflects_limit_of_reflects_isomorphisms _ _ + +instance forget_preserves_pullback_of_right : + preserves_limit (cospan g f) LocallyRingedSpace.forget_to_SheafedSpace := +preserves_pullback_symmetry _ _ _ + +instance forget_to_PresheafedSpace_preserves_pullback_of_right : + preserves_limit (cospan g f) (LocallyRingedSpace.forget_to_SheafedSpace ⋙ + SheafedSpace.forget_to_PresheafedSpace) := +preserves_pullback_symmetry _ _ _ + +instance forget_reflects_pullback_of_right : + reflects_limit (cospan g f) LocallyRingedSpace.forget_to_SheafedSpace := +reflects_limit_of_reflects_isomorphisms _ _ + +instance forget_to_PresheafedSpace_reflects_pullback_of_left : + reflects_limit (cospan f g) + (LocallyRingedSpace.forget_to_SheafedSpace ⋙ SheafedSpace.forget_to_PresheafedSpace) := +reflects_limit_of_reflects_isomorphisms _ _ + +instance forget_to_PresheafedSpace_reflects_pullback_of_right : + reflects_limit (cospan g f) + (LocallyRingedSpace.forget_to_SheafedSpace ⋙ SheafedSpace.forget_to_PresheafedSpace) := +reflects_limit_of_reflects_isomorphisms _ _ + +lemma pullback_snd_is_iso_of_range_subset (H' : set.range g.1.base ⊆ set.range f.1.base) : + is_iso (pullback.snd : pullback f g ⟶ _) := +begin + apply_with (reflects_isomorphisms.reflects LocallyRingedSpace.forget_to_SheafedSpace) + { instances := ff }, + apply_with (reflects_isomorphisms.reflects SheafedSpace.forget_to_PresheafedSpace) + { instances := ff }, + erw ← preserves_pullback.iso_hom_snd + (LocallyRingedSpace.forget_to_SheafedSpace ⋙ SheafedSpace.forget_to_PresheafedSpace) f g, + haveI := PresheafedSpace.is_open_immersion.pullback_snd_is_iso_of_range_subset _ _ H', + apply_instance, + apply_instance +end + +/-- +The universal property of open immersions: +For an open immersion `f : X ⟶ Z`, given any morphism of schemes `g : Y ⟶ Z` whose topological +image is contained in the image of `f`, we can lift this morphism to a unique `Y ⟶ X` that +commutes with these maps. +-/ +def lift (H' : set.range g.1.base ⊆ set.range f.1.base) : Y ⟶ X := +begin + haveI := pullback_snd_is_iso_of_range_subset f g H', + exact inv (pullback.snd : pullback f g ⟶ _) ≫ pullback.fst, +end + +@[simp, reassoc] lemma lift_fac (H' : set.range g.1.base ⊆ set.range f.1.base) : + lift f g H' ≫ f = g := +by { erw category.assoc, rw is_iso.inv_comp_eq, exact pullback.condition } + +lemma lift_uniq (H' : set.range g.1.base ⊆ set.range f.1.base) (l : Y ⟶ X) + (hl : l ≫ f = g) : l = lift f g H' := +by rw [← cancel_mono f, hl, lift_fac] + +lemma lift_range (H' : set.range g.1.base ⊆ set.range f.1.base) : + set.range (lift f g H').1.base = f.1.base ⁻¹' (set.range g.1.base) := +begin + haveI := pullback_snd_is_iso_of_range_subset f g H', + dsimp only [lift], + have : _ = (pullback.fst : pullback f g ⟶ _).val.base := preserves_pullback.iso_hom_fst + (LocallyRingedSpace.forget_to_SheafedSpace ⋙ SheafedSpace.forget _) f g, + rw [LocallyRingedSpace.comp_val, SheafedSpace.comp_base, ← this, ← category.assoc, coe_comp], + rw [set.range_comp, set.range_iff_surjective.mpr, set.image_univ, Top.pullback_fst_range], + ext, + split, + { rintros ⟨y, eq⟩, exact ⟨y, eq.symm⟩ }, + { rintros ⟨y, eq⟩, exact ⟨y, eq.symm⟩ }, + { rw ← Top.epi_iff_surjective, + rw (show (inv (pullback.snd : pullback f g ⟶ _)).val.base = _, from + (LocallyRingedSpace.forget_to_SheafedSpace ⋙ SheafedSpace.forget _).map_inv _), + apply_instance } +end + +end pullback + +/-- An open immersion is isomorphic to the induced open subscheme on its image. -/ +def iso_restrict {X Y : LocallyRingedSpace} {f : X ⟶ Y} + (H : LocallyRingedSpace.is_open_immersion f) : X ≅ Y.restrict H.base_open := +begin + apply LocallyRingedSpace.iso_of_SheafedSpace_iso, + refine SheafedSpace.forget_to_PresheafedSpace.preimage_iso _, + exact H.iso_restrict +end + +end LocallyRingedSpace.is_open_immersion + +section open_cover + +end open_cover + +end algebraic_geometry diff --git a/src/algebraic_geometry/presheafed_space/gluing.lean b/src/algebraic_geometry/presheafed_space/gluing.lean index df84093cc8460..8f214a1122052 100644 --- a/src/algebraic_geometry/presheafed_space/gluing.lean +++ b/src/algebraic_geometry/presheafed_space/gluing.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Andrew Yang -/ import topology.gluing -import algebraic_geometry.open_immersion +import algebraic_geometry.open_immersion.basic import algebraic_geometry.locally_ringed_space.has_colimits /-! From 308826471968962c6b59c7ff82a22757386603e3 Mon Sep 17 00:00:00 2001 From: damiano Date: Mon, 5 Jun 2023 02:20:49 +0000 Subject: [PATCH 1143/1291] chore(archive/imo): fix some naming inconsistencies and whitespace (#19155) I realize that this is the third PR concerning namespacing the files in `archive/imo`. These are some simple inconsistencies in naming. There is no rush in merging this PR, but in the long run it might help to have consistent namespacing, so that batch renames can be simpler. --- archive/imo/imo1975_q1.lean | 6 +----- archive/imo/imo1977_q6.lean | 1 + archive/imo/imo1988_q6.lean | 1 + archive/imo/imo1994_q1.lean | 1 + archive/imo/imo1998_q2.lean | 1 + archive/imo/imo2001_q2.lean | 1 + archive/imo/imo2005_q3.lean | 1 + archive/imo/imo2005_q4.lean | 8 +++++--- archive/imo/imo2006_q3.lean | 1 + archive/imo/imo2006_q5.lean | 1 + archive/imo/imo2008_q3.lean | 1 + archive/imo/imo2008_q4.lean | 1 + archive/imo/imo2013_q1.lean | 1 + archive/imo/imo2013_q5.lean | 1 + archive/imo/imo2019_q1.lean | 6 +----- archive/imo/imo2019_q2.lean | 1 + archive/imo/imo2019_q4.lean | 5 ++--- archive/imo/imo2021_q1.lean | 8 +++++--- 18 files changed, 27 insertions(+), 19 deletions(-) diff --git a/archive/imo/imo1975_q1.lean b/archive/imo/imo1975_q1.lean index 0035c3fe5d698..eaf260b9cdbac 100644 --- a/archive/imo/imo1975_q1.lean +++ b/archive/imo/imo1975_q1.lean @@ -31,9 +31,7 @@ variables (hx : antitone_on x (finset.Icc 1 n)) variables (hy : antitone_on y (finset.Icc 1 n)) include hx hy hσ -namespace imo1975_q1 - -theorem IMO_1975_Q1 : +theorem imo1975_q1 : ∑ i in finset.Icc 1 n, (x i - y i) ^ 2 ≤ ∑ i in finset.Icc 1 n, (x i - y (σ i)) ^ 2 := begin simp only [sub_sq, finset.sum_add_distrib, finset.sum_sub_distrib], @@ -48,5 +46,3 @@ begin -- them being `decreasing` exact antitone_on.monovary_on hx hy end - -end imo1975_q1 diff --git a/archive/imo/imo1977_q6.lean b/archive/imo/imo1977_q6.lean index 002795950f143..445644be7a341 100644 --- a/archive/imo/imo1977_q6.lean +++ b/archive/imo/imo1977_q6.lean @@ -36,6 +36,7 @@ begin end end imo1977_q6 + open imo1977_q6 theorem imo1977_q6 (f : ℕ+ → ℕ+) (h : ∀ n, f (f n) < f (n + 1)) : diff --git a/archive/imo/imo1988_q6.lean b/archive/imo/imo1988_q6.lean index 1f2888fbfa544..d5cd3fa2020e3 100644 --- a/archive/imo/imo1988_q6.lean +++ b/archive/imo/imo1988_q6.lean @@ -186,6 +186,7 @@ begin end end imo1988_q6 + open imo1988_q6 /--Question 6 of IMO1988. If a and b are two natural numbers diff --git a/archive/imo/imo1994_q1.lean b/archive/imo/imo1994_q1.lean index 7e6d6dc9b43cf..289fdf3250ad2 100644 --- a/archive/imo/imo1994_q1.lean +++ b/archive/imo/imo1994_q1.lean @@ -46,6 +46,7 @@ begin end end imo1994_q1 + open imo1994_q1 theorem imo1994_q1 (n : ℕ) (m : ℕ) (A : finset ℕ) (hm : A.card = m + 1) diff --git a/archive/imo/imo1998_q2.lean b/archive/imo/imo1998_q2.lean index 25557c40c8595..ab3cb7e1ccb35 100644 --- a/archive/imo/imo1998_q2.lean +++ b/archive/imo/imo1998_q2.lean @@ -198,6 +198,7 @@ lemma clear_denominators {a b k : ℕ} (ha : 0 < a) (hb : 0 < b) : by rw div_le_div_iff; norm_cast; simp [ha, hb] end imo1998_q2 + open imo1998_q2 theorem imo1998_q2 [fintype J] [fintype C] diff --git a/archive/imo/imo2001_q2.lean b/archive/imo/imo2001_q2.lean index 51ca0c9aec19e..fc2d5172f39e8 100644 --- a/archive/imo/imo2001_q2.lean +++ b/archive/imo/imo2001_q2.lean @@ -67,6 +67,7 @@ calc _ ≥ _ : add_le_add (add_le_add (bound ha hb hc) (bound hb hc ha)) (bound ... = 1 : by rw [h₁, h₂, ← add_div, ← add_div, div_self $ ne_of_gt $ denom_pos ha hb hc] end imo2001_q2 + open imo2001_q2 theorem imo2001_q2 (ha : 0 < a) (hb : 0 < b) (hc : 0 < c) : diff --git a/archive/imo/imo2005_q3.lean b/archive/imo/imo2005_q3.lean index 4cb5e03f1a9dd..bfcb287444ded 100644 --- a/archive/imo/imo2005_q3.lean +++ b/archive/imo/imo2005_q3.lean @@ -45,6 +45,7 @@ begin end end imo2005_q3 + open imo2005_q3 theorem imo2005_q3 (x y z : ℝ) (hx : x > 0) (hy : y > 0) (hz : z > 0) (h : x*y*z ≥ 1) : diff --git a/archive/imo/imo2005_q4.lean b/archive/imo/imo2005_q4.lean index b403688ea1702..9caceeddf5abf 100644 --- a/archive/imo/imo2005_q4.lean +++ b/archive/imo/imo2005_q4.lean @@ -55,8 +55,12 @@ begin ... = 0 : by ring, end +end imo2005_q4 + +open imo2005_q4 + /-- Main statement: The only positive integer coprime to all terms of the sequence `a` is `1`. -/ -example {k : ℕ} (hk : 0 < k) : (∀ n : ℕ, 1 ≤ n → is_coprime (a n) k) ↔ k = 1 := +theorem imo2005_q4 {k : ℕ} (hk : 0 < k) : (∀ n : ℕ, 1 ≤ n → is_coprime (a n) k) ↔ k = 1 := begin split, rotate, { -- The property is clearly true for `k = 1` @@ -95,5 +99,3 @@ begin -- Contradiction! contradiction, end - -end imo2005_q4 diff --git a/archive/imo/imo2006_q3.lean b/archive/imo/imo2006_q3.lean index 5d3378638a2e6..36a22e34bf2d0 100644 --- a/archive/imo/imo2006_q3.lean +++ b/archive/imo/imo2006_q3.lean @@ -136,6 +136,7 @@ begin end end imo2006_q3 + open imo2006_q3 theorem imo2006_q3 (M : ℝ) : diff --git a/archive/imo/imo2006_q5.lean b/archive/imo/imo2006_q5.lean index 1d0028ec3da52..682b1994adb37 100644 --- a/archive/imo/imo2006_q5.lean +++ b/archive/imo/imo2006_q5.lean @@ -201,6 +201,7 @@ begin end end imo2006_q5 + open imo2006_q5 /-- The general problem follows easily from the k = 2 case. -/ diff --git a/archive/imo/imo2008_q3.lean b/archive/imo/imo2008_q3.lean index 64e7080d53b99..16a937f43fce5 100644 --- a/archive/imo/imo2008_q3.lean +++ b/archive/imo/imo2008_q3.lean @@ -80,6 +80,7 @@ begin end end imo2008_q3 + open imo2008_q3 theorem imo2008_q3 : ∀ N : ℕ, ∃ n : ℕ, n ≥ N ∧ diff --git a/archive/imo/imo2008_q4.lean b/archive/imo/imo2008_q4.lean index 30d80008f57cb..70339914ed2d3 100644 --- a/archive/imo/imo2008_q4.lean +++ b/archive/imo/imo2008_q4.lean @@ -30,6 +30,7 @@ by rw [← pow_left_inj (abs_nonneg x) zero_le_one (pos_iff_ne_zero.2 hn), one_p abs_one] end imo2008_q4 + open imo2008_q4 theorem imo2008_q4 diff --git a/archive/imo/imo2013_q1.lean b/archive/imo/imo2013_q1.lean index 90b7943ee58ed..cf7c03ce84b00 100644 --- a/archive/imo/imo2013_q1.lean +++ b/archive/imo/imo2013_q1.lean @@ -46,6 +46,7 @@ begin end end imo2013_q1 + open imo2013_q1 theorem imo2013_q1 (n : ℕ+) (k : ℕ) : diff --git a/archive/imo/imo2013_q5.lean b/archive/imo/imo2013_q5.lean index 7211de3e55827..d30c32da64014 100644 --- a/archive/imo/imo2013_q5.lean +++ b/archive/imo/imo2013_q5.lean @@ -199,6 +199,7 @@ begin end end imo2013_q5 + open imo2013_q5 theorem imo2013_q5 diff --git a/archive/imo/imo2019_q1.lean b/archive/imo/imo2019_q1.lean index 6fd6d266a6a21..4e4772456a2ee 100644 --- a/archive/imo/imo2019_q1.lean +++ b/archive/imo/imo2019_q1.lean @@ -19,9 +19,7 @@ Note that there is a much more compact proof of this fact in Isabelle/HOL - http://downthetypehole.de/paste/4YbGgqb4 -/ -namespace imo2019_q1 - -theorem imo2019Q1 (f : ℤ → ℤ) : +theorem imo2019_q1 (f : ℤ → ℤ) : (∀ a b : ℤ, f (2 * a) + 2 * (f b) = f (f (a + b))) ↔ (f = 0) ∨ ∃ c, f = λ x, 2 * x + c := begin @@ -51,5 +49,3 @@ begin { right, use c, ext b, simp [H, add_comm] }, { left, ext b, simpa [H, two_ne_zero] using H3 } end - -end imo2019_q1 diff --git a/archive/imo/imo2019_q2.lean b/archive/imo/imo2019_q2.lean index 4289ec36cbd5a..7be26dade77e2 100644 --- a/archive/imo/imo2019_q2.lean +++ b/archive/imo/imo2019_q2.lean @@ -583,6 +583,7 @@ end end imo2019q2_cfg end imo2019_q2 + open imo2019_q2 theorem imo2019_q2 (A B C A₁ B₁ P Q P₁ Q₁ : Pt) diff --git a/archive/imo/imo2019_q4.lean b/archive/imo/imo2019_q4.lean index 479df77d88099..3465674493c60 100644 --- a/archive/imo/imo2019_q4.lean +++ b/archive/imo/imo2019_q4.lean @@ -28,7 +28,7 @@ open finset multiplicity nat (hiding zero_le prime) namespace imo2019_q4 -theorem imo2019_q4_upper_bound {k n : ℕ} (hk : k > 0) +theorem upper_bound {k n : ℕ} (hk : k > 0) (h : (k! : ℤ) = ∏ i in range n, (2 ^ n - 2 ^ i)) : n < 6 := begin have prime_2 : prime (2 : ℤ), @@ -70,7 +70,6 @@ begin end end imo2019_q4 -open imo2019_q4 theorem imo2019_q4 {k n : ℕ} (hk : k > 0) (hn : n > 0) : (k! : ℤ) = (∏ i in range n, (2 ^ n - 2 ^ i)) ↔ (k, n) = (1, 1) ∨ (k, n) = (3, 2) := @@ -81,7 +80,7 @@ begin norm_num [prod_range_succ, succ_mul] }, intro h, /- We know that n < 6. -/ - have := imo2019_q4_upper_bound hk h, + have := imo2019_q4.upper_bound hk h, interval_cases n, /- n = 1 -/ { left, congr, norm_num at h, rw [factorial_eq_one] at h, apply antisymm h, diff --git a/archive/imo/imo2021_q1.lean b/archive/imo/imo2021_q1.lean index 04659d2cb76a2..85abc0654bd36 100644 --- a/archive/imo/imo2021_q1.lean +++ b/archive/imo/imo2021_q1.lean @@ -152,7 +152,11 @@ begin rintros d (rfl|rfl|rfl); split; linarith only [hna, hab, hbc, hcn], }, end -theorem IMO_2021_Q1 : ∀ (n : ℕ), 100 ≤ n → ∀ (A ⊆ finset.Icc n (2 * n)), +end imo2021_q1 + +open imo2021_q1 + +theorem imo2021_q1 : ∀ (n : ℕ), 100 ≤ n → ∀ (A ⊆ finset.Icc n (2 * n)), (∃ (a b ∈ A), a ≠ b ∧ ∃ (k : ℕ), a + b = k * k) ∨ (∃ (a b ∈ finset.Icc n (2 * n) \ A), a ≠ b ∧ ∃ (k : ℕ), a + b = k * k) := begin @@ -180,5 +184,3 @@ begin cases hCA; [right, left]; exact ⟨a, (hCA ha).2, b, (hCA hb).2, hab, h₁ a (hCA ha).1 b (hCA hb).1 hab⟩, end - -end imo2021_q1 From df76f43357840485b9d04ed5dee5ab115d420e87 Mon Sep 17 00:00:00 2001 From: Riccardo Brasca Date: Mon, 5 Jun 2023 05:15:23 +0000 Subject: [PATCH 1144/1291] chore(field_theory/splitting_field): split file (#19154) We split `field_theory.splitting_field` into `field_theory.splitting_field.is_splitting_field` and `field_theory.splitting_field.construction`. This is useful for the port, but also quite a lot of Galois theory should not depend on the existence of splitting fields. --- src/field_theory/adjoin.lean | 2 +- src/field_theory/finite/galois_field.lean | 2 +- .../is_alg_closed/algebraic_closure.lean | 1 + src/field_theory/normal.lean | 10 +- src/field_theory/primitive_element.lean | 2 +- .../construction.lean} | 118 +------------- .../splitting_field/is_splitting_field.lean | 148 ++++++++++++++++++ src/ring_theory/polynomial/gauss_lemma.lean | 2 +- 8 files changed, 166 insertions(+), 119 deletions(-) rename src/field_theory/{splitting_field.lean => splitting_field/construction.lean} (79%) create mode 100644 src/field_theory/splitting_field/is_splitting_field.lean diff --git a/src/field_theory/adjoin.lean b/src/field_theory/adjoin.lean index 8d20609ac3665..13e44e04c536d 100644 --- a/src/field_theory/adjoin.lean +++ b/src/field_theory/adjoin.lean @@ -6,7 +6,7 @@ Authors: Thomas Browning, Patrick Lutz import field_theory.intermediate_field import field_theory.separable -import field_theory.splitting_field +import field_theory.splitting_field.is_splitting_field import ring_theory.tensor_product /-! diff --git a/src/field_theory/finite/galois_field.lean b/src/field_theory/finite/galois_field.lean index 9165891e346a9..bb674db9bbb99 100644 --- a/src/field_theory/finite/galois_field.lean +++ b/src/field_theory/finite/galois_field.lean @@ -8,7 +8,7 @@ import algebra.char_p.algebra import data.zmod.algebra import field_theory.finite.basic import field_theory.galois -import field_theory.splitting_field +import field_theory.splitting_field.is_splitting_field /-! # Galois fields diff --git a/src/field_theory/is_alg_closed/algebraic_closure.lean b/src/field_theory/is_alg_closed/algebraic_closure.lean index 18dc842228302..5f7ead381cc16 100644 --- a/src/field_theory/is_alg_closed/algebraic_closure.lean +++ b/src/field_theory/is_alg_closed/algebraic_closure.lean @@ -6,6 +6,7 @@ Authors: Kenny Lau import algebra.direct_limit import algebra.char_p.algebra import field_theory.is_alg_closed.basic +import field_theory.splitting_field.construction /-! # Algebraic Closure diff --git a/src/field_theory/normal.lean b/src/field_theory/normal.lean index 213fe17544515..c4b959da912e9 100644 --- a/src/field_theory/normal.lean +++ b/src/field_theory/normal.lean @@ -135,9 +135,11 @@ local attribute [-instance] adjoin_root.has_smul lemma normal.of_is_splitting_field (p : F[X]) [hFEp : is_splitting_field F E p] : normal F E := begin unfreezingI { rcases eq_or_ne p 0 with rfl | hp }, - { haveI : is_splitting_field F F 0 := ⟨splits_zero _, subsingleton.elim _ _⟩, - exact (alg_equiv.transfer_normal ((is_splitting_field.alg_equiv F 0).trans - (is_splitting_field.alg_equiv E 0).symm)).mp (normal_self F) }, + { have := hFEp.adjoin_roots, + simp only [polynomial.map_zero, roots_zero, multiset.to_finset_zero, finset.coe_empty, + algebra.adjoin_empty] at this, + exact normal.of_alg_equiv (alg_equiv.of_bijective (algebra.of_id F E) + (algebra.bijective_algebra_map_iff.2 this.symm)) }, refine normal_iff.2 (λ x, _), have hFE : finite_dimensional F E := is_splitting_field.finite_dimensional E p, have Hx : is_integral F x := is_integral_of_noetherian (is_noetherian.iff_fg.2 hFE) x, @@ -196,8 +198,6 @@ begin rw [set.image_singleton, ring_hom.algebra_map_to_algebra, adjoin_root.lift_root] end -instance (p : F[X]) : normal F p.splitting_field := normal.of_is_splitting_field p - end normal_tower namespace intermediate_field diff --git a/src/field_theory/primitive_element.lean b/src/field_theory/primitive_element.lean index 91f9463c8f69e..0b7587117823a 100644 --- a/src/field_theory/primitive_element.lean +++ b/src/field_theory/primitive_element.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Thomas Browning, Patrick Lutz -/ -import field_theory.adjoin +import field_theory.splitting_field.construction import field_theory.is_alg_closed.basic import field_theory.separable import ring_theory.integral_domain diff --git a/src/field_theory/splitting_field.lean b/src/field_theory/splitting_field/construction.lean similarity index 79% rename from src/field_theory/splitting_field.lean rename to src/field_theory/splitting_field/construction.lean index cf3ef1028d045..c160e90e30bc4 100644 --- a/src/field_theory/splitting_field.lean +++ b/src/field_theory/splitting_field/construction.lean @@ -3,29 +3,19 @@ Copyright (c) 2018 Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes -/ -import algebra.char_p.algebra -import field_theory.intermediate_field -import ring_theory.adjoin.field +import field_theory.normal /-! # Splitting fields -This file introduces the notion of a splitting field of a polynomial and provides an embedding from -a splitting field to any field that splits the polynomial. A polynomial `f : K[X]` splits -over a field extension `L` of `K` if it is zero or all of its irreducible factors over `L` have -degree `1`. A field extension of `K` of a polynomial `f : K[X]` is called a splitting field -if it is the smallest field extension of `K` such that `f` splits. +In this file we prove the existence and uniqueness of splitting fields. ## Main definitions * `polynomial.splitting_field f`: A fixed splitting field of the polynomial `f`. -* `polynomial.is_splitting_field`: A predicate on a field to be a splitting field of a polynomial - `f`. ## Main statements -* `polynomial.is_splitting_field.lift`: An embedding of a splitting field of the polynomial `f` into - another field such that `f` splits. * `polynomial.is_splitting_field.alg_equiv`: Every splitting field of a polynomial `f` is isomorphic to `splitting_field f` and thus, being a splitting field is unique up to isomorphism. @@ -484,81 +474,16 @@ adjoin_roots f end splitting_field -variables (K L) [algebra K L] -/-- Typeclass characterising splitting fields. -/ -class is_splitting_field (f : K[X]) : Prop := -(splits [] : splits (algebra_map K L) f) -(adjoin_roots [] : algebra.adjoin K (↑(f.map (algebra_map K L)).roots.to_finset : set L) = ⊤) +end splitting_field namespace is_splitting_field +variables (K L) [algebra K L] + variables {K} instance splitting_field (f : K[X]) : is_splitting_field K (splitting_field f) f := ⟨splitting_field.splits f, splitting_field.adjoin_roots f⟩ -section scalar_tower - -variables {K L F} [algebra F K] [algebra F L] [is_scalar_tower F K L] - -variables {K} -instance map (f : F[X]) [is_splitting_field F L f] : - is_splitting_field K L (f.map $ algebra_map F K) := -⟨by { rw [splits_map_iff, ← is_scalar_tower.algebra_map_eq], exact splits L f }, - subalgebra.restrict_scalars_injective F $ - by { rw [map_map, ← is_scalar_tower.algebra_map_eq, subalgebra.restrict_scalars_top, - eq_top_iff, ← adjoin_roots L f, algebra.adjoin_le_iff], - exact λ x hx, @algebra.subset_adjoin K _ _ _ _ _ _ hx }⟩ - -variables {K} (L) -theorem splits_iff (f : K[X]) [is_splitting_field K L f] : - polynomial.splits (ring_hom.id K) f ↔ (⊤ : subalgebra K L) = ⊥ := -⟨λ h, eq_bot_iff.2 $ adjoin_roots L f ▸ (roots_map (algebra_map K L) h).symm ▸ - algebra.adjoin_le_iff.2 (λ y hy, - let ⟨x, hxs, hxy⟩ := finset.mem_image.1 (by rwa multiset.to_finset_map at hy) in - hxy ▸ set_like.mem_coe.2 $ subalgebra.algebra_map_mem _ _), - λ h, @ring_equiv.to_ring_hom_refl K _ ▸ - ring_equiv.self_trans_symm (ring_equiv.of_bijective _ $ algebra.bijective_algebra_map_iff.2 h) ▸ - by { rw ring_equiv.to_ring_hom_trans, exact splits_comp_of_splits _ _ (splits L f) }⟩ - -theorem mul (f g : F[X]) (hf : f ≠ 0) (hg : g ≠ 0) [is_splitting_field F K f] - [is_splitting_field K L (g.map $ algebra_map F K)] : - is_splitting_field F L (f * g) := -⟨(is_scalar_tower.algebra_map_eq F K L).symm ▸ splits_mul _ - (splits_comp_of_splits _ _ (splits K f)) - ((splits_map_iff _ _).1 (splits L $ g.map $ algebra_map F K)), - by rw [polynomial.map_mul, roots_mul (mul_ne_zero (map_ne_zero hf : f.map (algebra_map F L) ≠ 0) - (map_ne_zero hg)), multiset.to_finset_add, finset.coe_union, - algebra.adjoin_union_eq_adjoin_adjoin, - is_scalar_tower.algebra_map_eq F K L, ← map_map, - roots_map (algebra_map K L) ((splits_id_iff_splits $ algebra_map F K).2 $ splits K f), - multiset.to_finset_map, finset.coe_image, algebra.adjoin_algebra_map, adjoin_roots, - algebra.map_top, is_scalar_tower.adjoin_range_to_alg_hom, ← map_map, adjoin_roots, - subalgebra.restrict_scalars_top]⟩ - -end scalar_tower - -/-- Splitting field of `f` embeds into any field that splits `f`. -/ -def lift [algebra K F] (f : K[X]) [is_splitting_field K L f] - (hf : polynomial.splits (algebra_map K F) f) : L →ₐ[K] F := -if hf0 : f = 0 then (algebra.of_id K F).comp $ - (algebra.bot_equiv K L : (⊥ : subalgebra K L) →ₐ[K] K).comp $ - by { rw ← (splits_iff L f).1 (show f.splits (ring_hom.id K), from hf0.symm ▸ splits_zero _), - exact algebra.to_top } else -alg_hom.comp (by { rw ← adjoin_roots L f, exact classical.choice (lift_of_splits _ $ λ y hy, - have aeval y f = 0, from (eval₂_eq_eval_map _).trans $ - (mem_roots $ by exact map_ne_zero hf0).1 (multiset.mem_to_finset.mp hy), - ⟨is_algebraic_iff_is_integral.1 ⟨f, hf0, this⟩, - splits_of_splits_of_dvd _ hf0 hf $ minpoly.dvd _ _ this⟩) }) - algebra.to_top - -theorem finite_dimensional (f : K[X]) [is_splitting_field K L f] : finite_dimensional K L := -⟨@algebra.top_to_submodule K L _ _ _ ▸ adjoin_roots L f ▸ - fg_adjoin_of_finite (finset.finite_to_set _) (λ y hy, - if hf : f = 0 - then by { rw [hf, polynomial.map_zero, roots_zero] at hy, cases hy } - else is_algebraic_iff_is_integral.1 ⟨f, hf, (eval₂_eq_eval_map _).trans $ - (mem_roots $ by exact map_ne_zero hf).1 (multiset.mem_to_finset.mp hy)⟩)⟩ - instance (f : K[X]) : _root_.finite_dimensional K f.splitting_field := finite_dimensional f.splitting_field f @@ -582,39 +507,12 @@ begin exact ring_hom.injective (lift L f $ splits (splitting_field f) f : L →+* f.splitting_field) end -lemma of_alg_equiv [algebra K F] (p : K[X]) (f : F ≃ₐ[K] L) [is_splitting_field K F p] : - is_splitting_field K L p := -begin - split, - { rw ← f.to_alg_hom.comp_algebra_map, - exact splits_comp_of_splits _ _ (splits F p) }, - { rw [←(algebra.range_top_iff_surjective f.to_alg_hom).mpr f.surjective, - ←root_set, adjoin_root_set_eq_range (splits F p), root_set, adjoin_roots F p] }, -end - end is_splitting_field -end splitting_field - end polynomial -namespace intermediate_field +section normal -open polynomial - -variables [field K] [field L] [algebra K L] {p : K[X]} - -lemma splits_of_splits {F : intermediate_field K L} (h : p.splits (algebra_map K L)) - (hF : ∀ x ∈ p.root_set L, x ∈ F) : p.splits (algebra_map K F) := -begin - simp_rw [root_set, finset.mem_coe, multiset.mem_to_finset] at hF, - rw splits_iff_exists_multiset, - refine ⟨multiset.pmap subtype.mk _ hF, map_injective _ (algebra_map F L).injective _⟩, - conv_lhs { rw [polynomial.map_map, ←is_scalar_tower.algebra_map_eq, - eq_prod_roots_of_splits h, ←multiset.pmap_eq_map _ _ _ hF] }, - simp_rw [polynomial.map_mul, polynomial.map_multiset_prod, - multiset.map_pmap, polynomial.map_sub, map_C, map_X], - refl, -end +instance [field F] (p : F[X]) : normal F p.splitting_field := normal.of_is_splitting_field p -end intermediate_field +end normal diff --git a/src/field_theory/splitting_field/is_splitting_field.lean b/src/field_theory/splitting_field/is_splitting_field.lean new file mode 100644 index 0000000000000..fd3dbd83e6d0d --- /dev/null +++ b/src/field_theory/splitting_field/is_splitting_field.lean @@ -0,0 +1,148 @@ +/- +Copyright (c) 2018 Chris Hughes. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Chris Hughes +-/ +import algebra.char_p.algebra +import field_theory.intermediate_field +import ring_theory.adjoin.field + +/-! +# Splitting fields + +This file introduces the notion of a splitting field of a polynomial and provides an embedding from +a splitting field to any field that splits the polynomial. A polynomial `f : K[X]` splits +over a field extension `L` of `K` if it is zero or all of its irreducible factors over `L` have +degree `1`. A field extension of `K` of a polynomial `f : K[X]` is called a splitting field +if it is the smallest field extension of `K` such that `f` splits. + +## Main definitions + +* `polynomial.is_splitting_field`: A predicate on a field to be a splitting field of a polynomial + `f`. + +## Main statements + +* `polynomial.is_splitting_field.lift`: An embedding of a splitting field of the polynomial `f` into + another field such that `f` splits. + +-/ + +noncomputable theory +open_locale classical big_operators polynomial + +universes u v w + +variables {F : Type u} (K : Type v) (L : Type w) + +namespace polynomial + +variables [field K] [field L] [field F] [algebra K L] + +/-- Typeclass characterising splitting fields. -/ +class is_splitting_field (f : K[X]) : Prop := +(splits [] : splits (algebra_map K L) f) +(adjoin_roots [] : algebra.adjoin K (↑(f.map (algebra_map K L)).roots.to_finset : set L) = ⊤) + +variables {K L F} + +namespace is_splitting_field + +section scalar_tower + +variables [algebra F K] [algebra F L] [is_scalar_tower F K L] + +instance map (f : F[X]) [is_splitting_field F L f] : + is_splitting_field K L (f.map $ algebra_map F K) := +⟨by { rw [splits_map_iff, ← is_scalar_tower.algebra_map_eq], exact splits L f }, + subalgebra.restrict_scalars_injective F $ + by { rw [map_map, ← is_scalar_tower.algebra_map_eq, subalgebra.restrict_scalars_top, + eq_top_iff, ← adjoin_roots L f, algebra.adjoin_le_iff], + exact λ x hx, @algebra.subset_adjoin K _ _ _ _ _ _ hx }⟩ + +variables (L) +theorem splits_iff (f : K[X]) [is_splitting_field K L f] : + polynomial.splits (ring_hom.id K) f ↔ (⊤ : subalgebra K L) = ⊥ := +⟨λ h, eq_bot_iff.2 $ adjoin_roots L f ▸ (roots_map (algebra_map K L) h).symm ▸ + algebra.adjoin_le_iff.2 (λ y hy, + let ⟨x, hxs, hxy⟩ := finset.mem_image.1 (by rwa multiset.to_finset_map at hy) in + hxy ▸ set_like.mem_coe.2 $ subalgebra.algebra_map_mem _ _), + λ h, @ring_equiv.to_ring_hom_refl K _ ▸ + ring_equiv.self_trans_symm (ring_equiv.of_bijective _ $ algebra.bijective_algebra_map_iff.2 h) ▸ + by { rw ring_equiv.to_ring_hom_trans, exact splits_comp_of_splits _ _ (splits L f) }⟩ + +theorem mul (f g : F[X]) (hf : f ≠ 0) (hg : g ≠ 0) [is_splitting_field F K f] + [is_splitting_field K L (g.map $ algebra_map F K)] : + is_splitting_field F L (f * g) := +⟨(is_scalar_tower.algebra_map_eq F K L).symm ▸ splits_mul _ + (splits_comp_of_splits _ _ (splits K f)) + ((splits_map_iff _ _).1 (splits L $ g.map $ algebra_map F K)), + by rw [polynomial.map_mul, roots_mul (mul_ne_zero (map_ne_zero hf : f.map (algebra_map F L) ≠ 0) + (map_ne_zero hg)), multiset.to_finset_add, finset.coe_union, + algebra.adjoin_union_eq_adjoin_adjoin, + is_scalar_tower.algebra_map_eq F K L, ← map_map, + roots_map (algebra_map K L) ((splits_id_iff_splits $ algebra_map F K).2 $ splits K f), + multiset.to_finset_map, finset.coe_image, algebra.adjoin_algebra_map, adjoin_roots, + algebra.map_top, is_scalar_tower.adjoin_range_to_alg_hom, ← map_map, adjoin_roots, + subalgebra.restrict_scalars_top]⟩ + +end scalar_tower + +variable (L) + +/-- Splitting field of `f` embeds into any field that splits `f`. -/ +def lift [algebra K F] (f : K[X]) [is_splitting_field K L f] + (hf : polynomial.splits (algebra_map K F) f) : L →ₐ[K] F := +if hf0 : f = 0 then (algebra.of_id K F).comp $ + (algebra.bot_equiv K L : (⊥ : subalgebra K L) →ₐ[K] K).comp $ + by { rw ← (splits_iff L f).1 (show f.splits (ring_hom.id K), from hf0.symm ▸ splits_zero _), + exact algebra.to_top } else +alg_hom.comp (by { rw ← adjoin_roots L f, exact classical.choice (lift_of_splits _ $ λ y hy, + have aeval y f = 0, from (eval₂_eq_eval_map _).trans $ + (mem_roots $ by exact map_ne_zero hf0).1 (multiset.mem_to_finset.mp hy), + ⟨is_algebraic_iff_is_integral.1 ⟨f, hf0, this⟩, + splits_of_splits_of_dvd _ hf0 hf $ minpoly.dvd _ _ this⟩) }) + algebra.to_top + +theorem finite_dimensional (f : K[X]) [is_splitting_field K L f] : finite_dimensional K L := +⟨@algebra.top_to_submodule K L _ _ _ ▸ adjoin_roots L f ▸ + fg_adjoin_of_finite (finset.finite_to_set _) (λ y hy, + if hf : f = 0 + then by { rw [hf, polynomial.map_zero, roots_zero] at hy, cases hy } + else is_algebraic_iff_is_integral.1 ⟨f, hf, (eval₂_eq_eval_map _).trans $ + (mem_roots $ by exact map_ne_zero hf).1 (multiset.mem_to_finset.mp hy)⟩)⟩ + +lemma of_alg_equiv [algebra K F] (p : K[X]) (f : F ≃ₐ[K] L) [is_splitting_field K F p] : + is_splitting_field K L p := +begin + split, + { rw ← f.to_alg_hom.comp_algebra_map, + exact splits_comp_of_splits _ _ (splits F p) }, + { rw [←(algebra.range_top_iff_surjective f.to_alg_hom).mpr f.surjective, + ←root_set, adjoin_root_set_eq_range (splits F p), root_set, adjoin_roots F p] }, +end + +end is_splitting_field + +end polynomial + +namespace intermediate_field + +open polynomial + +variables {K L} [field K] [field L] [algebra K L] {p : K[X]} + +lemma splits_of_splits {F : intermediate_field K L} (h : p.splits (algebra_map K L)) + (hF : ∀ x ∈ p.root_set L, x ∈ F) : p.splits (algebra_map K F) := +begin + simp_rw [root_set, finset.mem_coe, multiset.mem_to_finset] at hF, + rw splits_iff_exists_multiset, + refine ⟨multiset.pmap subtype.mk _ hF, map_injective _ (algebra_map F L).injective _⟩, + conv_lhs { rw [polynomial.map_map, ←is_scalar_tower.algebra_map_eq, + eq_prod_roots_of_splits h, ←multiset.pmap_eq_map _ _ _ hF] }, + simp_rw [polynomial.map_mul, polynomial.map_multiset_prod, + multiset.map_pmap, polynomial.map_sub, map_C, map_X], + refl, +end + +end intermediate_field diff --git a/src/ring_theory/polynomial/gauss_lemma.lean b/src/ring_theory/polynomial/gauss_lemma.lean index 33888b6a063ed..b4074084240df 100644 --- a/src/ring_theory/polynomial/gauss_lemma.lean +++ b/src/ring_theory/polynomial/gauss_lemma.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 Aaron Anderson. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Aaron Anderson -/ -import field_theory.splitting_field +import field_theory.splitting_field.construction import ring_theory.int.basic import ring_theory.localization.integral import ring_theory.integrally_closed From 4c3e1721c58ef9087bbc2c8c38b540f70eda2e53 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Roberto=20=C3=81lvarez?= Date: Mon, 5 Jun 2023 06:26:16 +0000 Subject: [PATCH 1145/1291] =?UTF-8?q?feat(topology/homotopy/homotopy=5Fgro?= =?UTF-8?q?up):=20`group`=20and=20`comm=5Fgroup`=20instances=20for=20`?= =?UTF-8?q?=CF=80=5Fn`=20(#15681)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR adds: - Group instance for `π_(n+1)` - Commutative group instance for `π_(n+2)` Co-authored-by: Junyan Xu Co-authored-by: Junyan Xu Co-authored-by: Kevin Buzzard --- src/topology/homeomorph.lean | 1 + src/topology/homotopy/homotopy_group.lean | 509 +++++++++++++++++----- src/topology/subset_properties.lean | 14 +- 3 files changed, 402 insertions(+), 122 deletions(-) diff --git a/src/topology/homeomorph.lean b/src/topology/homeomorph.lean index 1b72dea1ee01b..b57762b6e6c4e 100644 --- a/src/topology/homeomorph.lean +++ b/src/topology/homeomorph.lean @@ -504,6 +504,7 @@ variables [decidable_eq ι] (i : ι) continuous_inv_fun := continuous_pi $ λ j, by { dsimp only [equiv.pi_split_at], split_ifs, subst h, exacts [continuous_fst, (continuous_apply _).comp continuous_snd] } } +variable (β) /-- A product of copies of a topological space can be split as the binary product of one copy and the product of all the remaining copies. -/ @[simps] def fun_split_at : (ι → β) ≃ₜ β × ({j // j ≠ i} → β) := pi_split_at i _ diff --git a/src/topology/homotopy/homotopy_group.lean b/src/topology/homotopy/homotopy_group.lean index 8ca449bf538a4..021f8748bc561 100644 --- a/src/topology/homotopy/homotopy_group.lean +++ b/src/topology/homotopy/homotopy_group.lean @@ -5,6 +5,9 @@ Authors: Roberto Alvarez -/ import algebraic_topology.fundamental_groupoid.fundamental_group +import group_theory.eckmann_hilton +import logic.equiv.transfer_instance +import algebra.group.ext /-! # `n`th homotopy group @@ -12,187 +15,455 @@ import algebraic_topology.fundamental_groupoid.fundamental_group > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. > Any changes to this file require a corresponding PR to mathlib4. -We define the `n`th homotopy group at `x`, `π n x`, as the equivalence classes -of functions from the nth dimensional cube to the topological space `X` +We define the `n`th homotopy group at `x : X`, `π_n X x`, as the equivalence classes +of functions from the `n`-dimensional cube to the topological space `X` that send the boundary to the base point `x`, up to homotopic equivalence. -Note that such functions are generalized loops `gen_loop n x`, in particular -`gen_loop 1 x ≃ path x x` +Note that such functions are generalized loops `gen_loop (fin n) x`; in particular +`gen_loop (fin 1) x ≃ path x x`. -We show that `π 0 x` is equivalent to the path-conected components, and -that `π 1 x` is equivalent to the fundamental group at `x`. +We show that `π_0 X x` is equivalent to the path-connected components, and +that `π_1 X x` is equivalent to the fundamental group at `x`. +We provide a group instance using path composition and show commutativity when `n > 1`. ## definitions -* `gen_loop n x` is the type of continous fuctions `I^n → X` that send the boundary to `x` -* `homotopy_group n x` denoted `π n x` is the quotient of `gen_loop n x` by homotopy relative - to the boundary +* `gen_loop N x` is the type of continuous fuctions `I^N → X` that send the boundary to `x`, +* `homotopy_group.pi n X x` denoted `π_ n X x` is the quotient of `gen_loop (fin n) x` by + homotopy relative to the boundary, +* group instance `group (π_(n+1) X x)`, +* commutative group instance `comm_group (π_(n+2) X x)`. -TODO: show that `π n x` is a group when `n > 0` and abelian when `n > 1`. Show that -`pi1_equiv_fundamental_group` is an isomorphism of groups. +TODO: +* `Ω^M (Ω^N X) ≃ₜ Ω^(M⊕N) X`, and `Ω^M X ≃ₜ Ω^N X` when `M ≃ N`. Similarly for `π_`. +* Path-induced homomorphisms. Show that `homotopy_group.pi_1_equiv_fundamental_group` + is a group isomorphism. +* Examples with `𝕊^n`: `π_n (𝕊^n) = ℤ`, `π_m (𝕊^n)` trivial for `m < n`. +* Actions of π_1 on π_n. +* Lie algebra: `⁅π_(n+1), π_(m+1)⁆` contained in `π_(n+m+1)`. -/ open_locale unit_interval topology +open homeomorph noncomputable theory -universes u -variables {X : Type u} [topological_space X] -variables {n : ℕ} {x : X} - -/-- -The `n`-dimensional cube. --/ -@[derive [has_zero, has_one, topological_space]] -def cube (n : ℕ) : Type := fin n → I -local notation `I^` := cube +localized "notation `I^` N := N → I" in topology namespace cube -@[continuity] lemma proj_continuous (i : fin n) : continuous (λ f : I^n, f i) := -continuous_apply i +/-- The points in a cube with at least one projection equal to 0 or 1. -/ +def boundary (N : Type*) : set (I^N) := {y | ∃ i, y i = 0 ∨ y i = 1} -/-- -The points of the `n`-dimensional cube with at least one projection equal to 0 or 1. --/ -def boundary (n : ℕ) : set (I^n) := {y | ∃ i, y i = 0 ∨ y i = 1} +variables {N : Type*} [decidable_eq N] -/-- -The first projection of a positive-dimensional cube. --/ -@[simp] def head {n} : I^(n+1) → I := λ c, c 0 +/-- The forward direction of the homeomorphism + between the cube $I^N$ and $I × I^{N\setminus\{j\}}$. -/ +@[reducible] def split_at (i : N) : (I^N) ≃ₜ I × I^{j // j ≠ i} := fun_split_at I i -@[continuity] lemma head.continuous {n} : continuous (@head n) := proj_continuous 0 +/-- The backward direction of the homeomorphism + between the cube $I^N$ and $I × I^{N\setminus\{j\}}$. -/ +@[reducible] def insert_at (i : N) : I × (I^{j // j ≠ i}) ≃ₜ I^N := (fun_split_at I i).symm -/-- -The projection to the last `n` coordinates from an `n+1` dimensional cube. --/ -@[simp] def tail {n} : I^(n+1) → I^n := λ c, fin.tail c +lemma insert_at_boundary (i : N) {t₀ : I} {t} (H : (t₀ = 0 ∨ t₀ = 1) ∨ t ∈ boundary {j // j ≠ i}) : + insert_at i ⟨t₀, t⟩ ∈ boundary N := +begin + obtain H | ⟨j, H⟩ := H, + { use i, rwa [fun_split_at_symm_apply, dif_pos rfl] }, + { use j, rwa [fun_split_at_symm_apply, dif_neg j.prop, subtype.coe_eta] }, +end -instance unique_cube0 : unique (I^0) := pi.unique_of_is_empty _ +end cube -lemma one_char (f : I^1) : f = λ _, f 0 := by convert eq_const_of_unique f +variables (N X : Type*) [topological_space X] (x : X) -end cube +/-- The space of paths with both endpoints equal to a specified point `x : X`. -/ +@[reducible] def loop_space := path x x +localized "notation `Ω` := loop_space" in topology -/-- -The `n`-dimensional generalized loops; functions `I^n → X` that send the boundary to the base point. --/ -structure gen_loop (n : ℕ) (x : X) extends C(I^n, X) := -(boundary : ∀ y ∈ cube.boundary n, to_fun y = x) +instance loop_space.inhabited : inhabited (path x x) := ⟨path.refl x⟩ + +/-- The `n`-dimensional generalized loops based at `x` in a space `X` are + continuous functions `I^n → X` that sends the boundary to `x`. + We allow an arbitrary indexing type `N` in place of `fin n` here. -/ +def gen_loop : set C(I^N, X) := {p | ∀ y ∈ cube.boundary N, p y = x} +localized "notation `Ω^` := gen_loop" in topology + +variables {N X x} namespace gen_loop -instance fun_like : fun_like (gen_loop n x) (I^n) (λ _, X) := +/-- Copy of a `gen_loop` with a new map from the unit cube equal to the old one. + Useful to fix definitional equalities. -/ +def copy (f : Ω^N X x) (g : (I^N) → X) (h : g = f) : Ω^N X x := +⟨⟨g, h.symm ▸ f.1.2⟩, by { convert f.2, ext1, simp_rw h, refl }⟩ + +lemma coe_copy (f : Ω^N X x) {g : (I^N) → X} (h : g = f) : ⇑(copy f g h) = g := rfl + +lemma copy_eq (f : Ω^N X x) {g : (I^N) → X} (h : g = f) : copy f g h = f := +by { ext x, exact congr_fun h x } + +lemma boundary (f : Ω^N X x) : ∀ y ∈ cube.boundary N, f y = x := f.2 + +instance fun_like : fun_like (Ω^N X x) (I^N) (λ _, X) := { coe := λ f, f.1, coe_injective' := λ ⟨⟨f, _⟩, _⟩ ⟨⟨g, _⟩, _⟩ h, by { congr, exact h } } -@[ext] lemma ext (f g : gen_loop n x) (H : ∀ y, f y = g y) : f = g := fun_like.ext f g H +@[ext] lemma ext (f g : Ω^N X x) (H : ∀ y, f y = g y) : f = g := +fun_like.coe_injective' (funext H) -@[simp] lemma mk_apply (f : C(I^n, X)) (H y) : (⟨f, H⟩ : gen_loop n x) y = f y := rfl +@[simp] lemma mk_apply (f : C(I^N, X)) (H y) : (⟨f, H⟩ : Ω^N X x) y = f y := rfl -/-- -The constant `gen_loop` at `x`. --/ -def const : gen_loop n x := ⟨continuous_map.const _ x, λ _ _, rfl⟩ +/-- The constant `gen_loop` at `x`. -/ +def const : Ω^N X x := ⟨continuous_map.const _ x, λ _ _, rfl⟩ -instance inhabited : inhabited (gen_loop n x) := { default := const } +@[simp] lemma const_apply {t} : (@const N X _ x) t = x := rfl -/-- -The "homotopy relative to boundary" relation between `gen_loop`s. --/ -def homotopic (f g : gen_loop n x) : Prop := -f.to_continuous_map.homotopic_rel g.to_continuous_map (cube.boundary n) +instance inhabited : inhabited (Ω^N X x) := ⟨const⟩ + +/-- The "homotopic relative to boundary" relation between `gen_loop`s. -/ +def homotopic (f g : Ω^N X x) : Prop := f.1.homotopic_rel g.1 (cube.boundary N) namespace homotopic -section -variables {f g h : gen_loop n x} -@[refl] lemma refl (f : gen_loop n x) : homotopic f f := continuous_map.homotopic_rel.refl _ +variables {f g h : Ω^N X x} + +@[refl] lemma refl (f : Ω^N X x) : homotopic f f := continuous_map.homotopic_rel.refl _ -@[symm] lemma symm (H : f.homotopic g) : g.homotopic f := H.symm +@[symm] lemma symm (H : homotopic f g) : homotopic g f := H.symm -@[trans] lemma trans (H0 : f.homotopic g) (H1 : g.homotopic h) : f.homotopic h := H0.trans H1 +@[trans] lemma trans (H0 : homotopic f g) (H1 : homotopic g h) : homotopic f h := H0.trans H1 -lemma equiv : equivalence (@homotopic X _ n x) := +lemma equiv : equivalence (@homotopic N X _ x) := ⟨homotopic.refl, λ _ _, homotopic.symm, λ _ _ _, homotopic.trans⟩ -instance setoid (n : ℕ) (x : X) : setoid (gen_loop n x) := ⟨homotopic, equiv⟩ +instance setoid (N) (x : X) : setoid (Ω^N X x) := ⟨homotopic, equiv⟩ -end end homotopic +section loop_homeo + +variable [decidable_eq N] + +/-- Loop from a generalized loop by currying $I^N → X$ into $I → (I^{N\setminus\{j\}} → X)$. -/ +@[simps] def to_loop (i : N) (p : Ω^N X x) : Ω (Ω^{j // j ≠ i} X x) const := +{ to_fun := λ t, ⟨(p.val.comp (cube.insert_at i).to_continuous_map).curry t, + λ y yH, p.property (cube.insert_at i (t, y)) (cube.insert_at_boundary i $ or.inr yH)⟩, + source' := by { ext t, refine p.property (cube.insert_at i (0, t)) ⟨i, or.inl _⟩, simp }, + target' := by { ext t, refine p.property (cube.insert_at i (1, t)) ⟨i, or.inr _⟩, simp } } + +lemma continuous_to_loop (i : N) : continuous (@to_loop N X _ x _ i) := +path.continuous_uncurry_iff.1 $ continuous.subtype_mk (continuous_map.continuous_eval'.comp $ + continuous.prod_map (continuous_map.continuous_curry.comp $ + (continuous_map.continuous_comp_left _).comp continuous_subtype_coe) continuous_id) _ + +/-- Generalized loop from a loop by uncurrying $I → (I^{N\setminus\{j\}} → X)$ into $I^N → X$. -/ +@[simps] def from_loop (i : N) (p : Ω (Ω^{j // j ≠ i} X x) const) : Ω^N X x := +⟨(continuous_map.comp ⟨coe⟩ p.to_continuous_map).uncurry.comp (cube.split_at i).to_continuous_map, +begin + rintros y ⟨j, Hj⟩, + simp only [subtype.val_eq_coe, continuous_map.comp_apply, to_continuous_map_apply, + fun_split_at_apply, continuous_map.uncurry_apply, continuous_map.coe_mk, + function.uncurry_apply_pair], + obtain rfl | Hne := eq_or_ne j i, + { cases Hj; simpa only [Hj, p.coe_to_continuous_map, p.source, p.target] }, + { exact gen_loop.boundary _ _ ⟨⟨j, Hne⟩, Hj⟩ }, +end⟩ + +lemma continuous_from_loop (i : N) : continuous (@from_loop N X _ x _ i) := +((continuous_map.continuous_comp_left _).comp $ continuous_map.continuous_uncurry.comp $ + (continuous_map.continuous_comp _).comp continuous_induced_dom).subtype_mk _ + +lemma to_from (i : N) (p : Ω (Ω^{j // j ≠ i} X x) const) : to_loop i (from_loop i p) = p := +begin + simp_rw [to_loop, from_loop, continuous_map.comp_assoc, to_continuous_map_as_coe, + to_continuous_map_comp_symm, continuous_map.comp_id], ext, refl, +end + +/-- The `n+1`-dimensional loops are in bijection with the loops in the space of + `n`-dimensional loops with base point `const`. + We allow an arbitrary indexing type `N` in place of `fin n` here. -/ +@[simps] def loop_homeo (i : N) : Ω^N X x ≃ₜ Ω (Ω^{j // j ≠ i} X x) const := +{ to_fun := to_loop i, + inv_fun := from_loop i, + left_inv := λ p, by { ext, exact congr_arg p (equiv.apply_symm_apply _ _) }, + right_inv := to_from i, + continuous_to_fun := continuous_to_loop i, + continuous_inv_fun := continuous_from_loop i } + +lemma to_loop_apply (i : N) {p : Ω^N X x} {t} {tn} : + to_loop i p t tn = p (cube.insert_at i ⟨t, tn⟩) := rfl + +lemma from_loop_apply (i : N) {p : Ω (Ω^{j // j ≠ i} X x) const} {t : I^N} : + from_loop i p t = p (t i) (cube.split_at i t).snd := rfl + +/-- Composition with `cube.insert_at` as a continuous map. -/ +@[reducible] def c_comp_insert (i : N) : C(C(I^N, X), C(I × I^{j // j ≠ i}, X)) := +⟨λ f, f.comp (cube.insert_at i).to_continuous_map, + (cube.insert_at i).to_continuous_map.continuous_comp_left⟩ + +/-- A homotopy between `n+1`-dimensional loops `p` and `q` constant on the boundary + seen as a homotopy between two paths in the space of `n`-dimensional paths. -/ +def homotopy_to (i : N) {p q : Ω^N X x} (H : p.1.homotopy_rel q.1 (cube.boundary N)) : + C(I × I, C(I^{j // j ≠ i}, X)) := +((⟨_, continuous_map.continuous_curry⟩: C(_,_)).comp $ + (c_comp_insert i).comp H.to_continuous_map.curry).uncurry + +-- Should be generated with `@[simps]` but it times out. +lemma homotopy_to_apply (i : N) {p q : Ω^N X x} (H : p.1.homotopy_rel q.1 $ cube.boundary N) + (t : I × I) (tₙ : I^{j // j ≠ i}) : + homotopy_to i H t tₙ = H (t.fst, cube.insert_at i (t.snd, tₙ)) := rfl + +lemma homotopic_to (i : N) {p q : Ω^N X x} : + homotopic p q → (to_loop i p).homotopic (to_loop i q) := +begin + refine nonempty.map (λ H, ⟨⟨⟨λ t, ⟨homotopy_to i H t, _⟩, _⟩, _, _⟩, _⟩), + { rintros y ⟨i, iH⟩, + rw [homotopy_to_apply, H.eq_fst, p.2], + all_goals { apply cube.insert_at_boundary, right, exact ⟨i, iH⟩} }, + { continuity }, + show ∀ _ _ _, _, + { intros t y yH, + split; ext; erw homotopy_to_apply, + apply H.eq_fst, work_on_goal 2 { apply H.eq_snd }, + all_goals { use i, rw [fun_split_at_symm_apply, dif_pos rfl], exact yH } }, + all_goals { intro, ext, erw [homotopy_to_apply, to_loop_apply] }, + exacts [H.apply_zero _, H.apply_one _], +end + +/-- The converse to `gen_loop.homotopy_to`: a homotopy between two loops in the space of + `n`-dimensional loops can be seen as a homotopy between two `n+1`-dimensional paths. -/ +def homotopy_from (i : N) {p q : Ω^N X x} + (H : (to_loop i p).homotopy (to_loop i q)) : C(I × I^N, X) := +(continuous_map.comp ⟨_, continuous_map.continuous_uncurry⟩ + (continuous_map.comp ⟨coe⟩ H.to_continuous_map).curry).uncurry.comp $ + (continuous_map.id I).prod_map (cube.split_at i).to_continuous_map + +-- Should be generated with `@[simps]` but it times out. +lemma homotopy_from_apply (i : N) {p q : Ω^N X x} (H : (to_loop i p).homotopy (to_loop i q)) + (t : I × I^N) : homotopy_from i H t = H (t.fst, t.snd i) (λ j, t.snd ↑j) := rfl + +lemma homotopic_from (i : N) {p q : Ω^N X x} : + (to_loop i p).homotopic (to_loop i q) → homotopic p q := +begin + refine nonempty.map (λ H, ⟨⟨homotopy_from i H, _, _⟩, _⟩), + show ∀ _ _ _, _, + { rintros t y ⟨j, jH⟩, + erw homotopy_from_apply, + obtain rfl | h := eq_or_ne j i, + { split, + { rw H.eq_fst, exacts [congr_arg p (equiv.right_inv _ _), jH] }, + { rw H.eq_snd, exacts [congr_arg q (equiv.right_inv _ _), jH] } }, + { rw [p.2 _ ⟨j, jH⟩, q.2 _ ⟨j, jH⟩], split; { apply boundary, exact ⟨⟨j, h⟩, jH⟩ } } }, + all_goals { intro, + convert homotopy_from_apply _ _ _, + rw H.apply_zero <|> rw H.apply_one, + apply congr_arg p <|> apply congr_arg q, + exact (equiv.right_inv _ _).symm }, +end + +/-- Concatenation of two `gen_loop`s along the `i`th coordinate. -/ +def trans_at (i : N) (f g : Ω^N X x) : Ω^N X x := +copy (from_loop i $ (to_loop i f).trans $ to_loop i g) + (λ t, if (t i : ℝ) ≤ 1/2 + then f (t.update i $ set.proj_Icc 0 1 zero_le_one (2 * t i)) + else g (t.update i $ set.proj_Icc 0 1 zero_le_one (2 * t i - 1))) +begin + ext1, symmetry, + dsimp only [path.trans, from_loop, path.coe_mk, function.comp_app, + mk_apply, continuous_map.comp_apply, to_continuous_map_apply, fun_split_at_apply, + continuous_map.uncurry_apply, continuous_map.coe_mk, function.uncurry_apply_pair], + split_ifs, change f _ = _, swap, change g _ = _, + all_goals { congr' 1 } +end + +/-- Reversal of a `gen_loop` along the `i`th coordinate. -/ +def symm_at (i : N) (f : Ω^N X x) : Ω^N X x := +copy (from_loop i (to_loop i f).symm) + (λ t, f $ λ j, if j = i then σ (t i) else t j) $ + by { ext1, change _ = f _, congr, ext1, simp } + +lemma trans_at_distrib {i j : N} (h : i ≠ j) (a b c d : Ω^N X x) : + trans_at i (trans_at j a b) (trans_at j c d) = trans_at j (trans_at i a c) (trans_at i b d) := +begin + ext, simp_rw [trans_at, coe_copy, function.update_apply, if_neg h, if_neg h.symm], + split_ifs; { congr' 1, ext1, simp only [function.update, eq_rec_constant, dite_eq_ite], + apply ite_ite_comm, rintro rfl, exact h.symm }, +end + +lemma from_loop_trans_to_loop {i : N} {p q : Ω^N X x} : + from_loop i ((to_loop i p).trans $ to_loop i q) = trans_at i p q := +(copy_eq _ _).symm + +lemma from_loop_symm_to_loop {i : N} {p : Ω^N X x} : + from_loop i (to_loop i p).symm = symm_at i p := (copy_eq _ _).symm + +end loop_homeo + end gen_loop -/-- -The `n`th homotopy group at `x` defined as the quotient of `gen_loop n x` by the -`homotopic` relation. --/ +/-- The `n`th homotopy group at `x` defined as the quotient of `Ω^n x` by the + `gen_loop.homotopic` relation. -/ @[derive inhabited] -def homotopy_group (n : ℕ) (x : X) : Type _ := quotient (gen_loop.homotopic.setoid n x) -local notation `π` := homotopy_group +def homotopy_group (N) (X : Type*) [topological_space X] (x : X) : Type* := +quotient (gen_loop.homotopic.setoid N x) + +variable [decidable_eq N] +open gen_loop +/-- Equivalence between the homotopy group of X and the fundamental group of + `Ω^{j // j ≠ i} x`. -/ +def homotopy_group_equiv_fundamental_group (i : N) : + homotopy_group N X x ≃ fundamental_group (Ω^{j // j ≠ i} X x) const := +begin + refine equiv.trans _ (category_theory.groupoid.iso_equiv_hom _ _).symm, + apply quotient.congr (loop_homeo i).to_equiv, + exact λ p q, ⟨homotopic_to i, homotopic_from i⟩, +end -/-- The 0-dimensional generalized loops based at `x` are in 1-1 correspondence with `X`. -/ -def gen_loop_zero_equiv : gen_loop 0 x ≃ X := -{ to_fun := λ f, f 0, - inv_fun := λ x, ⟨continuous_map.const _ x, λ _ ⟨f0,_⟩, f0.elim0⟩, - left_inv := λ f, by { ext1, exact congr_arg f (subsingleton.elim _ _) }, - right_inv := λ _, rfl } +/-- Homotopy group of finite index. -/ +@[reducible] def homotopy_group.pi (n) (X : Type*) [topological_space X] (x : X) := +homotopy_group (fin n) _ x +localized "notation `π_` := homotopy_group.pi" in topology -/-- -The 0th homotopy "group" is equivalent to the path components of `X`, aka the `zeroth_homotopy`. --/ -def pi0_equiv_path_components : π 0 x ≃ zeroth_homotopy X := -quotient.congr gen_loop_zero_equiv +/-- The 0-dimensional generalized loops based at `x` are in bijection with `X`. -/ +def gen_loop_homeo_of_is_empty (N x) [is_empty N] : Ω^N X x ≃ₜ X := +{ to_fun := λ f, f 0, + inv_fun := λ y, ⟨continuous_map.const _ y, λ _ ⟨i, _⟩, is_empty_elim i⟩, + left_inv := λ f, by { ext, exact congr_arg f (subsingleton.elim _ _) }, + right_inv := λ _, rfl, + continuous_to_fun := + (continuous_map.continuous_eval_const' (0 : N → I)).comp continuous_induced_dom, + continuous_inv_fun := (continuous_map.const'.2).subtype_mk _ } + +/-- The homotopy "group" indexed by an empty type is in bijection with + the path components of `X`, aka the `zeroth_homotopy`. -/ +def homotopy_group_equiv_zeroth_homotopy_of_is_empty (N x) [is_empty N] : + homotopy_group N X x ≃ zeroth_homotopy X := +quotient.congr (gen_loop_homeo_of_is_empty N x).to_equiv begin -- joined iff homotopic intros, split; rintro ⟨H⟩, exacts - [⟨{ to_fun := λ t, H ⟨t, fin.elim0⟩, - source' := (H.apply_zero _).trans (congr_arg a₁ matrix.zero_empty.symm), - target' := (H.apply_one _).trans (congr_arg a₂ matrix.zero_empty.symm) }⟩, + [⟨{ to_fun := λ t, H ⟨t, is_empty_elim⟩, + source' := (H.apply_zero _).trans (congr_arg a₁ $ subsingleton.elim _ _), + target' := (H.apply_one _).trans (congr_arg a₂ $ subsingleton.elim _ _) }⟩, ⟨{ to_fun := λ t0, H t0.fst, map_zero_left' := λ _, by convert H.source, map_one_left' := λ _, by convert H.target, - prop' := λ _ _ ⟨i,_⟩, i.elim0 }⟩] + prop' := λ _ _ ⟨i, _⟩, is_empty_elim i }⟩], end -/-- The 1-dimensional generalized loops based at `x` are in 1-1 correspondence with - paths from `x` to itself. -/ -@[simps] def gen_loop_one_equiv_path_self : gen_loop 1 x ≃ path x x := -{ to_fun := λ p, path.mk ⟨λ t, p (λ _, t), by {continuity, exact p.1.2}⟩ - (p.boundary (λ _, 0) ⟨0, or.inl rfl⟩) - (p.boundary (λ _, 1) ⟨1, or.inr rfl⟩), - inv_fun := λ p, - { to_fun := λ c, p c.head, - boundary := begin - rintro y ⟨i, iH|iH⟩; cases unique.eq_default i; - apply (congr_arg p iH).trans, exacts [p.source, p.target], - end }, - left_inv := λ p, by { ext1, exact congr_arg p y.one_char.symm }, +/-- The 0th homotopy "group" is in bijection with `zeroth_homotopy`. -/ +def homotopy_group.pi_0_equiv_zeroth_homotopy : π_ 0 X x ≃ zeroth_homotopy X := +homotopy_group_equiv_zeroth_homotopy_of_is_empty (fin 0) x + +/-- The 1-dimensional generalized loops based at `x` are in bijection with loops at `x`. -/ +@[simps] def gen_loop_equiv_of_unique (N) [unique N] : Ω^N X x ≃ Ω X x := +{ to_fun := λ p, path.mk ⟨λ t, p (λ _, t), by continuity⟩ + (gen_loop.boundary _ (λ _, 0) ⟨default, or.inl rfl⟩) + (gen_loop.boundary _ (λ _, 1) ⟨default, or.inr rfl⟩), + inv_fun := λ p, ⟨⟨λ c, p (c default), by continuity⟩, + begin + rintro y ⟨i, iH|iH⟩; cases unique.eq_default i; apply (congr_arg p iH).trans, + exacts [p.source, p.target], + end⟩, + left_inv := λ p, by { ext, exact congr_arg p (eq_const_of_unique y).symm }, right_inv := λ p, by { ext, refl } } -/-- -The first homotopy group at `x` is equivalent to the fundamental group, -i.e. the loops based at `x` up to homotopy. --/ -def pi1_equiv_fundamental_group : π 1 x ≃ fundamental_group X x := +/-- The homotopy group at `x` indexed by a singleton is in bijection with the fundamental group, + i.e. the loops based at `x` up to homotopy. -/ +/- TODO (?): deducing this from `homotopy_group_equiv_fundamental_group` would require + combination of `category_theory.functor.map_Aut` and + `fundamental_groupoid.fundamental_groupoid_functor` applied to `gen_loop_homeo_of_is_empty`, + with possibly worse defeq. -/ +def homotopy_group_equiv_fundamental_group_of_unique (N) [unique N] : + homotopy_group N X x ≃ fundamental_group X x := begin refine equiv.trans _ (category_theory.groupoid.iso_equiv_hom _ _).symm, - refine quotient.congr gen_loop_one_equiv_path_self _, - -- homotopic iff homotopic + refine quotient.congr (gen_loop_equiv_of_unique N) _, intros, split; rintros ⟨H⟩, - exacts - [⟨{ to_fun := λ tx, H (tx.fst, λ _, tx.snd), - map_zero_left' := λ _, by convert H.apply_zero _, - map_one_left' := λ _, by convert H.apply_one _, - prop' := λ t y iH, H.prop' _ _ ⟨0,iH⟩ }⟩, - ⟨{ to_fun := λ tx, H (tx.fst, tx.snd.head), - map_zero_left' := λ y, by { convert H.apply_zero _, exact y.one_char }, - map_one_left' := λ y, by { convert H.apply_one _, exact y.one_char }, - prop' := λ t y ⟨i, iH⟩, begin - cases unique.eq_default i, split, - { convert H.eq_fst _ _, exacts [y.one_char, iH] }, - { convert H.eq_snd _ _, exacts [y.one_char, iH] }, - end }⟩], + { exact + ⟨ { to_fun := λ tx, H (tx.fst, λ _, tx.snd), + map_zero_left' := λ _, H.apply_zero _, + map_one_left' := λ _, H.apply_one _, + prop' := λ t y iH, H.prop' _ _ ⟨default, iH⟩ } ⟩ }, + refine ⟨⟨⟨⟨λ tx, H (tx.fst, tx.snd default), H.continuous.comp _⟩, λ y, _, λ y, _⟩, _⟩⟩, + { exact continuous_fst.prod_mk ((continuous_apply _).comp continuous_snd) }, + { convert H.apply_zero _, exact eq_const_of_unique y }, + { convert H.apply_one _, exact eq_const_of_unique y }, + { rintro t y ⟨i, iH⟩, + cases unique.eq_default i, split, + { convert H.eq_fst _ _, exacts [eq_const_of_unique y, iH] }, + { convert H.eq_snd _ _, exacts [eq_const_of_unique y, iH] } }, +end + +/-- The first homotopy group at `x` is in bijection with the fundamental group. -/ +def homotopy_group.pi_1_equiv_fundamental_group : π_ 1 X x ≃ fundamental_group X x := +homotopy_group_equiv_fundamental_group_of_unique (fin 1) + +namespace homotopy_group + +/-- Group structure on `homotopy_group N X x` for nonempty `N` (in particular `π_(n+1) X x`). -/ +instance group (N) [decidable_eq N] [nonempty N] : group (homotopy_group N X x) := +(homotopy_group_equiv_fundamental_group $ classical.arbitrary N).group + +/-- Group structure on `homotopy_group` obtained by pulling back path composition along the + `i`th direction. The group structures for two different `i j : N` distribute over each + other, and therefore are equal by the Eckmann-Hilton argument. -/ +@[reducible] def aux_group (i : N) : group (homotopy_group N X x) := +(homotopy_group_equiv_fundamental_group i).group + +lemma is_unital_aux_group (i : N) : + eckmann_hilton.is_unital (aux_group i).mul (⟦const⟧ : homotopy_group N X x) := +⟨⟨(aux_group i).one_mul⟩, ⟨(aux_group i).mul_one⟩⟩ + +lemma aux_group_indep (i j : N) : (aux_group i : group (homotopy_group N X x)) = aux_group j := +begin + by_cases h : i = j, { rw h }, + refine group.ext (eckmann_hilton.mul (is_unital_aux_group i) (is_unital_aux_group j) _), + rintro ⟨a⟩ ⟨b⟩ ⟨c⟩ ⟨d⟩, + change quotient.mk _ = _, + apply congr_arg quotient.mk, + simp only [from_loop_trans_to_loop, trans_at_distrib h, + coe_to_equiv, loop_homeo_apply, coe_symm_to_equiv, loop_homeo_symm_apply], end + +lemma trans_at_indep {i} (j) (f g : Ω^N X x) : ⟦trans_at i f g⟧ = ⟦trans_at j f g⟧ := +begin + simp_rw ← from_loop_trans_to_loop, + have := congr_arg (@group.mul _) (aux_group_indep i j), + exact congr_fun₂ this ⟦g⟧ ⟦f⟧, +end + +lemma symm_at_indep {i} (j) (f : Ω^N X x) : ⟦symm_at i f⟧ = ⟦symm_at j f⟧ := +begin + simp_rw ← from_loop_symm_to_loop, + have := congr_arg (@group.inv _) (aux_group_indep i j), + exact congr_fun this ⟦f⟧, +end + +/-- Characterization of multiplicative identity -/ +lemma one_def [nonempty N] : (1 : homotopy_group N X x) = ⟦const⟧ := rfl + +/-- Characterization of multiplication -/ +lemma mul_spec [nonempty N] {i} {p q : Ω^N X x} : + (⟦p⟧ * ⟦q⟧ : homotopy_group N X x) = ⟦trans_at i q p⟧ := +by { rw [trans_at_indep _ q, ← from_loop_trans_to_loop], apply quotient.sound, refl } + +/-- Characterization of multiplicative inverse -/ +lemma inv_spec [nonempty N] {i} {p : Ω^N X x} : (⟦p⟧⁻¹ : homotopy_group N X x) = ⟦symm_at i p⟧ := +by { rw [symm_at_indep _ p, ← from_loop_symm_to_loop], apply quotient.sound, refl } + +/-- Multiplication on `homotopy_group N X x` is commutative for nontrivial `N`. + In particular, multiplication on `π_(n+2)` is commutative. -/ +instance comm_group [nontrivial N] : comm_group (homotopy_group N X x) := +let h := exists_ne (classical.arbitrary N) in +@eckmann_hilton.comm_group (homotopy_group N X x) _ 1 (is_unital_aux_group h.some) _ +begin + rintro ⟨a⟩ ⟨b⟩ ⟨c⟩ ⟨d⟩, + apply congr_arg quotient.mk, + simp only [from_loop_trans_to_loop, trans_at_distrib h.some_spec, + coe_to_equiv, loop_homeo_apply, coe_symm_to_equiv, loop_homeo_symm_apply], +end + +end homotopy_group diff --git a/src/topology/subset_properties.lean b/src/topology/subset_properties.lean index b7da3d32318db..36f63e1120bbb 100644 --- a/src/topology/subset_properties.lean +++ b/src/topology/subset_properties.lean @@ -1050,6 +1050,8 @@ by { convert is_compact_pi_infinite h, simp only [← mem_univ_pi, set_of_mem_eq instance pi.compact_space [∀ i, compact_space (π i)] : compact_space (Πi, π i) := ⟨by { rw [← pi_univ univ], exact is_compact_univ_pi (λ i, is_compact_univ) }⟩ +instance function.compact_space [compact_space β] : compact_space (ι → β) := pi.compact_space + /-- **Tychonoff's theorem** formulated in terms of filters: `filter.cocompact` on an indexed product type `Π d, κ d` the `filter.Coprod` of filters `filter.cocompact` on `κ d`. -/ lemma filter.Coprod_cocompact {δ : Type*} {κ : δ → Type*} [Π d, topological_space (κ d)] : @@ -1093,7 +1095,7 @@ lemma locally_compact_space_of_has_basis {ι : α → Type*} {p : Π x, ι x → locally_compact_space α := ⟨λ x t ht, let ⟨i, hp, ht⟩ := (h x).mem_iff.1 ht in ⟨s x i, (h x).mem_of_mem hp, ht, hc x i hp⟩⟩ -instance locally_compact_space.prod (α : Type*) (β : Type*) [topological_space α] +instance prod.locally_compact_space (α : Type*) (β : Type*) [topological_space α] [topological_space β] [locally_compact_space α] [locally_compact_space β] : locally_compact_space (α × β) := have _ := λ x : α × β, (compact_basis_nhds x.1).prod_nhds' (compact_basis_nhds x.2), @@ -1105,7 +1107,7 @@ variables [Π i, topological_space (π i)] [∀ i, locally_compact_space (π i)] /--In general it suffices that all but finitely many of the spaces are compact, but that's not straightforward to state and use. -/ -instance locally_compact_space.pi_finite [finite ι] : locally_compact_space (Π i, π i) := +instance pi.locally_compact_space_of_finite [finite ι] : locally_compact_space (Π i, π i) := ⟨λ t n hn, begin rw [nhds_pi, filter.mem_pi] at hn, obtain ⟨s, hs, n', hn', hsub⟩ := hn, @@ -1116,7 +1118,7 @@ instance locally_compact_space.pi_finite [finite ι] : locally_compact_space (Π end⟩ /-- For spaces that are not Hausdorff. -/ -instance locally_compact_space.pi [∀ i, compact_space (π i)] : locally_compact_space (Π i, π i) := +instance pi.locally_compact_space [∀ i, compact_space (π i)] : locally_compact_space (Π i, π i) := ⟨λ t n hn, begin rw [nhds_pi, filter.mem_pi] at hn, obtain ⟨s, hs, n', hn', hsub⟩ := hn, @@ -1131,6 +1133,12 @@ instance locally_compact_space.pi [∀ i, compact_space (π i)] : locally_compac { rw if_neg h, exact compact_space.is_compact_univ, } }, end⟩ +instance function.locally_compact_space_of_finite [finite ι] [locally_compact_space β] : + locally_compact_space (ι → β) := pi.locally_compact_space_of_finite + +instance function.locally_compact_space [locally_compact_space β] [compact_space β] : + locally_compact_space (ι → β) := pi.locally_compact_space + end pi /-- A reformulation of the definition of locally compact space: In a locally compact space, From 575b4ea3738b017e30fb205cb9b4a8742e5e82b6 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Mon, 5 Jun 2023 06:26:18 +0000 Subject: [PATCH 1146/1291] chore(*): add mathlib4 synchronization comments (#19157) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.lie.tensor_product` * `analysis.calculus.cont_diff` * `analysis.calculus.inverse` * `analysis.calculus.lagrange_multipliers` * `analysis.complex.real_deriv` * `analysis.special_functions.exp_deriv` * `analysis.special_functions.sqrt` * `category_theory.monoidal.functor_category` * `linear_algebra.tensor_power` * `number_theory.sum_four_squares` --- src/algebra/lie/tensor_product.lean | 3 +++ src/analysis/calculus/cont_diff.lean | 3 +++ src/analysis/calculus/inverse.lean | 3 +++ src/analysis/calculus/lagrange_multipliers.lean | 3 +++ src/analysis/complex/real_deriv.lean | 3 +++ src/analysis/special_functions/exp_deriv.lean | 3 +++ src/analysis/special_functions/sqrt.lean | 3 +++ src/category_theory/monoidal/functor_category.lean | 3 +++ src/linear_algebra/tensor_power.lean | 3 +++ src/number_theory/sum_four_squares.lean | 3 +++ 10 files changed, 30 insertions(+) diff --git a/src/algebra/lie/tensor_product.lean b/src/algebra/lie/tensor_product.lean index ed0389173dcbc..1512f182e6fbc 100644 --- a/src/algebra/lie/tensor_product.lean +++ b/src/algebra/lie/tensor_product.lean @@ -8,6 +8,9 @@ import algebra.lie.abelian /-! # Tensor products of Lie modules +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Tensor products of Lie modules carry natural Lie module structures. ## Tags diff --git a/src/analysis/calculus/cont_diff.lean b/src/analysis/calculus/cont_diff.lean index 21556cd0fda2b..6df49e6855c02 100644 --- a/src/analysis/calculus/cont_diff.lean +++ b/src/analysis/calculus/cont_diff.lean @@ -12,6 +12,9 @@ import data.nat.choose.cast /-! # Higher differentiability of usual operations +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We prove that the usual operations (addition, multiplication, difference, composition, and so on) preserve `C^n` functions. We also expand the API around `C^n` functions. diff --git a/src/analysis/calculus/inverse.lean b/src/analysis/calculus/inverse.lean index de850dabe3232..f84292e53e633 100644 --- a/src/analysis/calculus/inverse.lean +++ b/src/analysis/calculus/inverse.lean @@ -9,6 +9,9 @@ import analysis.normed_space.banach /-! # Inverse function theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove the inverse function theorem. It says that if a map `f : E → F` has an invertible strict derivative `f'` at `a`, then it is locally invertible, and the inverse function has derivative `f' ⁻¹`. diff --git a/src/analysis/calculus/lagrange_multipliers.lean b/src/analysis/calculus/lagrange_multipliers.lean index 98273ea1853e9..4304f3d350b42 100644 --- a/src/analysis/calculus/lagrange_multipliers.lean +++ b/src/analysis/calculus/lagrange_multipliers.lean @@ -9,6 +9,9 @@ import linear_algebra.dual /-! # Lagrange multipliers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we formalize the [Lagrange multipliers](https://en.wikipedia.org/wiki/Lagrange_multiplier) method of solving conditional extremum problems: if a function `φ` has a local extremum at `x₀` on the set diff --git a/src/analysis/complex/real_deriv.lean b/src/analysis/complex/real_deriv.lean index b88e9735e27cc..b5923fc628e07 100644 --- a/src/analysis/complex/real_deriv.lean +++ b/src/analysis/complex/real_deriv.lean @@ -10,6 +10,9 @@ import analysis.calculus.conformal.normed_space /-! # Real differentiability of complex-differentiable functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + `has_deriv_at.real_of_complex` expresses that, if a function on `ℂ` is differentiable (over `ℂ`), then its restriction to `ℝ` is differentiable over `ℝ`, with derivative the real part of the complex derivative. diff --git a/src/analysis/special_functions/exp_deriv.lean b/src/analysis/special_functions/exp_deriv.lean index 666baa6e15813..828a3c2d14373 100644 --- a/src/analysis/special_functions/exp_deriv.lean +++ b/src/analysis/special_functions/exp_deriv.lean @@ -8,6 +8,9 @@ import analysis.complex.real_deriv /-! # Complex and real exponential +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove that `complex.exp` and `real.exp` are infinitely smooth functions. ## Tags diff --git a/src/analysis/special_functions/sqrt.lean b/src/analysis/special_functions/sqrt.lean index 3288842d449ea..4d47dc7972a8b 100644 --- a/src/analysis/special_functions/sqrt.lean +++ b/src/analysis/special_functions/sqrt.lean @@ -8,6 +8,9 @@ import analysis.calculus.cont_diff /-! # Smoothness of `real.sqrt` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove that `real.sqrt` is infinitely smooth at all points `x ≠ 0` and provide some dot-notation lemmas. diff --git a/src/category_theory/monoidal/functor_category.lean b/src/category_theory/monoidal/functor_category.lean index 0f5bc2d773ff6..937c29de0c9b9 100644 --- a/src/category_theory/monoidal/functor_category.lean +++ b/src/category_theory/monoidal/functor_category.lean @@ -10,6 +10,9 @@ import category_theory.functor.const /-! # Monoidal structure on `C ⥤ D` when `D` is monoidal. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + When `C` is any category, and `D` is a monoidal category, there is a natural "pointwise" monoidal structure on `C ⥤ D`. diff --git a/src/linear_algebra/tensor_power.lean b/src/linear_algebra/tensor_power.lean index a17d3b28dab14..94287845afd8e 100644 --- a/src/linear_algebra/tensor_power.lean +++ b/src/linear_algebra/tensor_power.lean @@ -11,6 +11,9 @@ import algebra.direct_sum.algebra /-! # Tensor power of a semimodule over a commutative semirings +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define the `n`th tensor power of `M` as the n-ary tensor product indexed by `fin n` of `M`, `⨂[R] (i : fin n), M`. This is a special case of `pi_tensor_product`. diff --git a/src/number_theory/sum_four_squares.lean b/src/number_theory/sum_four_squares.lean index c8aee31197186..8ea1a82447bc6 100644 --- a/src/number_theory/sum_four_squares.lean +++ b/src/number_theory/sum_four_squares.lean @@ -13,6 +13,9 @@ import data.fintype.big_operators /-! # Lagrange's four square theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The main result in this file is `sum_four_squares`, a proof that every natural number is the sum of four square numbers. From a3209ddf94136d36e5e5c624b10b2a347cc9d090 Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Mon, 5 Jun 2023 08:24:03 +0000 Subject: [PATCH 1147/1291] feat (special_functions/gamma): doubling formula (#19136) This PR adds the proof of the Legendre doubling formula, relating `Gamma s * Gamma (s + 1 / 2)` to `Gamma (2 * s)`. --- src/analysis/analytic/isolated_zeros.lean | 11 ++ src/analysis/analytic/uniqueness.lean | 12 +- .../special_functions/gamma/beta.lean | 68 +++++++- .../gamma/bohr_mollerup.lean | 159 ++++++++++++++++-- 4 files changed, 235 insertions(+), 15 deletions(-) diff --git a/src/analysis/analytic/isolated_zeros.lean b/src/analysis/analytic/isolated_zeros.lean index 8d7b7cffa7e1b..98343d71e3c9a 100644 --- a/src/analysis/analytic/isolated_zeros.lean +++ b/src/analysis/analytic/isolated_zeros.lean @@ -195,4 +195,15 @@ theorem eq_on_of_preconnected_of_mem_closure (hf : analytic_on 𝕜 f U) (hg : a eq_on f g U := hf.eq_on_of_preconnected_of_frequently_eq hg hU h₀ (mem_closure_ne_iff_frequently_within.mp hfg) +/-- The *identity principle* for analytic functions, global version: if two functions on a normed +field `𝕜` are analytic everywhere and coincide at points which accumulate to a point `z₀`, then +they coincide globally. +For higher-dimensional versions requiring that the functions coincide in a neighborhood of `z₀`, +see `eq_of_eventually_eq`. -/ +theorem eq_of_frequently_eq [connected_space 𝕜] + (hf : analytic_on 𝕜 f univ) (hg : analytic_on 𝕜 g univ) + (hfg : ∃ᶠ z in 𝓝[≠] z₀, f z = g z) : f = g := +funext (λ x, eq_on_of_preconnected_of_frequently_eq hf hg is_preconnected_univ + (mem_univ z₀) hfg (mem_univ x)) + end analytic_on diff --git a/src/analysis/analytic/uniqueness.lean b/src/analysis/analytic/uniqueness.lean index 57a5268a74db5..93dd668cc9339 100644 --- a/src/analysis/analytic/uniqueness.lean +++ b/src/analysis/analytic/uniqueness.lean @@ -89,7 +89,7 @@ begin exact uniform_space.completion.coe_injective F this, end -/-- The *identity principle* for analytic functions: If two analytic function coincide in a whole +/-- The *identity principle* for analytic functions: If two analytic functions coincide in a whole neighborhood of a point `z₀`, then they coincide globally along a connected set. For a one-dimensional version assuming only that the functions coincide at some points arbitrarily close to `z₀`, see `eq_on_of_preconnected_of_frequently_eq`. -/ @@ -103,4 +103,14 @@ begin λ z hz, (hf.sub hg).eq_on_zero_of_preconnected_of_eventually_eq_zero hU h₀ hfg' hz, end +/-- The *identity principle* for analytic functions: If two analytic functions on a normed space +coincide in a neighborhood of a point `z₀`, then they coincide everywhere. +For a one-dimensional version assuming only that the functions coincide at some points +arbitrarily close to `z₀`, see `eq_of_frequently_eq`. -/ +theorem eq_of_eventually_eq {f g : E → F} [preconnected_space E] + (hf : analytic_on 𝕜 f univ) (hg : analytic_on 𝕜 g univ) {z₀ : E} (hfg : f =ᶠ[𝓝 z₀] g) : + f = g := +funext (λ x, eq_on_of_preconnected_of_eventually_eq hf hg is_preconnected_univ + (mem_univ z₀) hfg (mem_univ x)) + end analytic_on diff --git a/src/analysis/special_functions/gamma/beta.lean b/src/analysis/special_functions/gamma/beta.lean index 6c24104b458b7..42a05012901ec 100644 --- a/src/analysis/special_functions/gamma/beta.lean +++ b/src/analysis/special_functions/gamma/beta.lean @@ -5,7 +5,8 @@ Authors: David Loeffler -/ import analysis.convolution import analysis.special_functions.trigonometric.euler_sine_prod -import analysis.special_functions.gamma.basic +import analysis.special_functions.gamma.bohr_mollerup +import analysis.analytic.isolated_zeros /-! # The Beta function, and further properties of the Gamma function @@ -28,8 +29,10 @@ refined properties of the Gamma function using these relations. * `complex.Gamma_mul_Gamma_one_sub`: Euler's reflection formula `Gamma s * Gamma (1 - s) = π / sin π s`. * `complex.differentiable_one_div_Gamma`: the function `1 / Γ(s)` is differentiable everywhere. +* `complex.Gamma_mul_Gamma_add_half`: Legendre's duplication formula + `Gamma s * Gamma (s + 1 / 2) = Gamma (2 * s) * 2 ^ (1 - 2 * s) * sqrt π`. * `real.Gamma_ne_zero`, `real.Gamma_seq_tendsto_Gamma`, - `real.Gamma_mul_Gamma_one_sub`: real versions of the above results. + `real.Gamma_mul_Gamma_one_sub`, `real.Gamma_mul_Gamma_add_half`: real versions of the above. -/ noncomputable theory @@ -560,3 +563,64 @@ end end complex end inv_gamma + +section doubling +/-! +## The doubling formula for Gamma + +We prove the doubling formula for arbitrary real or complex arguments, by analytic continuation from +the positive real case. (Knowing that `Γ⁻¹` is analytic everywhere makes this much simpler, since we +do not have to do any special-case handling for the poles of `Γ`.) +-/ + +namespace complex + +theorem Gamma_mul_Gamma_add_half (s : ℂ) : + Gamma s * Gamma (s + 1 / 2) = Gamma (2 * s) * 2 ^ (1 - 2 * s) * ↑(real.sqrt π) := +begin + suffices : (λ z, (Gamma z)⁻¹ * (Gamma (z + 1 / 2))⁻¹) = + (λ z, (Gamma (2 * z))⁻¹ * 2 ^ (2 * z - 1) / ↑(real.sqrt π)), + { convert congr_arg has_inv.inv (congr_fun this s) using 1, + { rw [mul_inv, inv_inv, inv_inv] }, + { rw [div_eq_mul_inv, mul_inv, mul_inv, inv_inv, inv_inv, ←cpow_neg, neg_sub] } }, + have h1 : analytic_on ℂ (λ z : ℂ, (Gamma z)⁻¹ * (Gamma (z + 1 / 2))⁻¹) univ, + { refine differentiable_on.analytic_on _ is_open_univ, + refine (differentiable_one_div_Gamma.mul _).differentiable_on, + exact differentiable_one_div_Gamma.comp (differentiable_id.add (differentiable_const _)) }, + have h2 : analytic_on ℂ (λ z, (Gamma (2 * z))⁻¹ * 2 ^ (2 * z - 1) / ↑(real.sqrt π)) univ, + { refine differentiable_on.analytic_on _ is_open_univ, + refine (differentiable.mul _ (differentiable_const _)).differentiable_on, + apply differentiable.mul, + { exact differentiable_one_div_Gamma.comp (differentiable_id'.const_mul _) }, + { refine λ t, differentiable_at.const_cpow _ (or.inl two_ne_zero), + refine differentiable_at.sub_const (differentiable_at_id.const_mul _) _ } }, + have h3 : tendsto (coe : ℝ → ℂ) (𝓝[≠] 1) (𝓝[≠] 1), + { rw tendsto_nhds_within_iff, split, + { exact tendsto_nhds_within_of_tendsto_nhds continuous_of_real.continuous_at }, + { exact eventually_nhds_within_iff.mpr (eventually_of_forall $ λ t ht, of_real_ne_one.mpr ht)}}, + refine analytic_on.eq_of_frequently_eq h1 h2 (h3.frequently _), + refine ((eventually.filter_mono nhds_within_le_nhds) _).frequently, + refine (eventually_gt_nhds zero_lt_one).mp (eventually_of_forall $ λ t ht, _), + rw [←mul_inv, Gamma_of_real, (by push_cast : (t : ℂ) + 1 / 2 = ↑(t + 1 / 2)), Gamma_of_real, + ←of_real_mul, Gamma_mul_Gamma_add_half_of_pos ht, of_real_mul, of_real_mul, ←Gamma_of_real, + mul_inv, mul_inv, (by push_cast : 2 * (t : ℂ) = ↑(2 * t)), Gamma_of_real, + of_real_cpow zero_le_two, of_real_bit0, of_real_one, ←cpow_neg, of_real_sub, of_real_one, + neg_sub, ←div_eq_mul_inv] +end + +end complex + +namespace real +open complex + +lemma Gamma_mul_Gamma_add_half (s : ℝ) : + Gamma s * Gamma (s + 1 / 2) = Gamma (2 * s) * 2 ^ (1 - 2 * s) * sqrt π := +begin + rw [←of_real_inj], + simpa only [←Gamma_of_real, of_real_cpow zero_le_two, of_real_mul, of_real_add, of_real_div, + of_real_bit0, of_real_one, of_real_sub] using complex.Gamma_mul_Gamma_add_half ↑s +end + +end real + +end doubling diff --git a/src/analysis/special_functions/gamma/bohr_mollerup.lean b/src/analysis/special_functions/gamma/bohr_mollerup.lean index b6aa7efc4ecbb..9db12f986223c 100644 --- a/src/analysis/special_functions/gamma/bohr_mollerup.lean +++ b/src/analysis/special_functions/gamma/bohr_mollerup.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: David Loeffler -/ import analysis.special_functions.gamma.basic +import analysis.special_functions.gaussian /-! # Convexity properties of the Gamma function @@ -25,11 +26,65 @@ context of this proof, we place them in a separate namespace `real.bohr_mollerup general form of the Euler limit formula valid for any real or complex `x`; see `real.Gamma_seq_tendsto_Gamma` and `complex.Gamma_seq_tendsto_Gamma` in the file `analysis.special_functions.gamma.beta`.) + +As an application of the Bohr-Mollerup theorem we prove the Legendre doubling formula for the +Gamma function for real positive `s` (which will be upgraded to a proof for all complex `s` in a +later file). + +TODO: This argument can be extended to prove the general `k`-multiplication formula (at least up +to a constant, and it should be possible to deduce the value of this constant using Stirling's +formula). -/ noncomputable theory open filter set measure_theory -open_locale nat ennreal topology big_operators +open_locale nat ennreal topology big_operators real + +section convexity + +-- Porting note: move the following lemmas to `Analysis.Convex.Function` +variables {𝕜 E β : Type*} {s : set E} {f g : E → β} + [ordered_semiring 𝕜] [has_smul 𝕜 E] [add_comm_monoid E] [ordered_add_comm_monoid β] + +lemma convex_on.congr [has_smul 𝕜 β] (hf : convex_on 𝕜 s f) (hfg : eq_on f g s) : + convex_on 𝕜 s g := +⟨hf.1, λ x hx y hy a b ha hb hab, + by simpa only [←hfg hx, ←hfg hy, ←hfg (hf.1 hx hy ha hb hab)] using hf.2 hx hy ha hb hab⟩ + +lemma concave_on.congr [has_smul 𝕜 β](hf : concave_on 𝕜 s f) (hfg : eq_on f g s) : + concave_on 𝕜 s g := +⟨hf.1, λ x hx y hy a b ha hb hab, + by simpa only [←hfg hx, ←hfg hy, ←hfg (hf.1 hx hy ha hb hab)] using hf.2 hx hy ha hb hab⟩ + +lemma strict_convex_on.congr [has_smul 𝕜 β] (hf : strict_convex_on 𝕜 s f) (hfg : eq_on f g s) : + strict_convex_on 𝕜 s g := +⟨hf.1, λ x hx y hy hxy a b ha hb hab, by simpa only + [←hfg hx, ←hfg hy, ←hfg (hf.1 hx hy ha.le hb.le hab)] using hf.2 hx hy hxy ha hb hab⟩ + +lemma strict_concave_on.congr [has_smul 𝕜 β] (hf : strict_concave_on 𝕜 s f) (hfg : eq_on f g s) : + strict_concave_on 𝕜 s g := +⟨hf.1, λ x hx y hy hxy a b ha hb hab, by simpa only + [←hfg hx, ←hfg hy, ←hfg (hf.1 hx hy ha.le hb.le hab)] using hf.2 hx hy hxy ha hb hab⟩ + +lemma convex_on.add_const [module 𝕜 β] (hf : convex_on 𝕜 s f) (b : β) : + convex_on 𝕜 s (f + (λ _, b)) := +hf.add (convex_on_const _ hf.1) + +lemma concave_on.add_const [module 𝕜 β] (hf : concave_on 𝕜 s f) (b : β) : + concave_on 𝕜 s (f + (λ _, b)) := +hf.add (concave_on_const _ hf.1) + +lemma strict_convex_on.add_const {γ : Type*} {f : E → γ} + [ordered_cancel_add_comm_monoid γ] [module 𝕜 γ] (hf : strict_convex_on 𝕜 s f) (b : γ) : + strict_convex_on 𝕜 s (f + (λ _, b)) := +hf.add_convex_on (convex_on_const _ hf.1) + +lemma strict_concave_on.add_const {γ : Type*} {f : E → γ} + [ordered_cancel_add_comm_monoid γ] [module 𝕜 γ] (hf : strict_concave_on 𝕜 s f) (b : γ) : + strict_concave_on 𝕜 s (f + (λ _, b)) := +hf.add_concave_on (concave_on_const _ hf.1) + +end convexity namespace real @@ -112,17 +167,12 @@ end lemma convex_on_Gamma : convex_on ℝ (Ioi 0) Gamma := begin - refine ⟨convex_Ioi 0, λ x hx y hy a b ha hb hab, _⟩, - have := convex_on.comp (convex_on_exp.subset (subset_univ _) _) convex_on_log_Gamma - (λ u hu v hv huv, exp_le_exp.mpr huv), - convert this.2 hx hy ha hb hab, - { rw [function.comp_app, exp_log (Gamma_pos_of_pos $ this.1 hx hy ha hb hab)] }, - { rw [function.comp_app, exp_log (Gamma_pos_of_pos hx)] }, - { rw [function.comp_app, exp_log (Gamma_pos_of_pos hy)] }, - { rw convex_iff_is_preconnected, - refine is_preconnected_Ioi.image _ (λ x hx, continuous_at.continuous_within_at _), - refine (differentiable_at_Gamma (λ m, _)).continuous_at.log (Gamma_pos_of_pos hx).ne', - exact (neg_lt_iff_pos_add.mpr (add_pos_of_pos_of_nonneg hx (nat.cast_nonneg m))).ne' } + refine ((convex_on_exp.subset (subset_univ _) _).comp convex_on_log_Gamma + (exp_monotone.monotone_on _)).congr (λ x hx, exp_log (Gamma_pos_of_pos hx)), + rw convex_iff_is_preconnected, + refine is_preconnected_Ioi.image _ (λ x hx, continuous_at.continuous_within_at _), + refine (differentiable_at_Gamma (λ m, _)).continuous_at.log (Gamma_pos_of_pos hx).ne', + exact (neg_lt_iff_pos_add.mpr (add_pos_of_pos_of_nonneg hx (nat.cast_nonneg m))).ne', end end convexity @@ -377,4 +427,89 @@ end end strict_mono +section doubling + +/-! +## The Gamma doubling formula + +As a fun application of the Bohr-Mollerup theorem, we prove the Gamma-function doubling formula +(for positive real `s`). The idea is that `2 ^ s * Gamma (s / 2) * Gamma (s / 2 + 1 / 2)` is +log-convex and satisfies the Gamma functional equation, so it must actually be a constant +multiple of `Gamma`, and we can compute the constant by specialising at `s = 1`. -/ + +/-- Auxiliary definition for the doubling formula (we'll show this is equal to `Gamma s`) -/ +def doubling_Gamma (s : ℝ) : ℝ := Gamma (s / 2) * Gamma (s / 2 + 1 / 2) * 2 ^ (s - 1) / sqrt π + +lemma doubling_Gamma_add_one (s : ℝ) (hs : s ≠ 0) : + doubling_Gamma (s + 1) = s * doubling_Gamma s := +begin + rw [doubling_Gamma, doubling_Gamma, (by abel : s + 1 - 1 = s - 1 + 1), add_div, add_assoc, + add_halves (1 : ℝ), Gamma_add_one (div_ne_zero hs two_ne_zero), rpow_add two_pos, rpow_one], + ring, +end + +lemma doubling_Gamma_one : doubling_Gamma 1 = 1 := +by simp_rw [doubling_Gamma, Gamma_one_half_eq, add_halves (1 : ℝ), sub_self, Gamma_one, mul_one, + rpow_zero, mul_one, div_self (sqrt_ne_zero'.mpr pi_pos)] + +lemma log_doubling_Gamma_eq : + eq_on (log ∘ doubling_Gamma) (λ s, log (Gamma (s / 2)) + log (Gamma (s / 2 + 1 / 2)) + + s * log 2 - log (2 * sqrt π)) (Ioi 0) := +begin + intros s hs, + have h1 : sqrt π ≠ 0, from sqrt_ne_zero'.mpr pi_pos, + have h2 : Gamma (s / 2) ≠ 0, from (Gamma_pos_of_pos $ div_pos hs two_pos).ne', + have h3 : Gamma (s / 2 + 1 / 2) ≠ 0, + from (Gamma_pos_of_pos $ add_pos (div_pos hs two_pos) one_half_pos).ne', + have h4 : (2 : ℝ) ^ (s - 1) ≠ 0, from (rpow_pos_of_pos two_pos _).ne', + rw [function.comp_app, doubling_Gamma, log_div (mul_ne_zero (mul_ne_zero h2 h3) h4) h1, + log_mul (mul_ne_zero h2 h3) h4, log_mul h2 h3, log_rpow two_pos, log_mul two_ne_zero h1], + ring_nf, +end + +lemma doubling_Gamma_log_convex_Ioi : convex_on ℝ (Ioi (0:ℝ)) (log ∘ doubling_Gamma) := +begin + refine (((convex_on.add _ _).add _).add_const _).congr log_doubling_Gamma_eq.symm, + { convert convex_on_log_Gamma.comp_affine_map + (distrib_mul_action.to_linear_map ℝ ℝ (1 / 2 : ℝ)).to_affine_map, + { simpa only [zero_div] using (preimage_const_mul_Ioi (0 : ℝ) one_half_pos).symm, }, + { ext1 x, + change log (Gamma (x / 2)) = log (Gamma ((1 / 2 : ℝ) • x)), + rw [smul_eq_mul, mul_comm, mul_one_div] } }, + { refine convex_on.subset _ (Ioi_subset_Ioi $ neg_one_lt_zero.le) (convex_Ioi _), + convert convex_on_log_Gamma.comp_affine_map ((distrib_mul_action.to_linear_map ℝ ℝ + (1 / 2 : ℝ)).to_affine_map + affine_map.const _ _ (1 / 2 : ℝ)), + { change Ioi (-1 : ℝ) = ((λ x : ℝ, x + 1 / 2) ∘ (λ x : ℝ, (1 / 2 : ℝ) * x)) ⁻¹' (Ioi 0), + rw [preimage_comp, preimage_add_const_Ioi, zero_sub, preimage_const_mul_Ioi (_ : ℝ) + one_half_pos, neg_div, div_self (@one_half_pos ℝ _).ne'] }, + { ext1 x, + change log (Gamma (x / 2 + 1 / 2)) = log (Gamma ((1 / 2 : ℝ) • x + 1 / 2)), + rw [smul_eq_mul, mul_comm, mul_one_div] } }, + { simpa only [mul_comm _ (log _)] + using (convex_on_id (convex_Ioi (0 : ℝ))).smul (log_pos one_lt_two).le } +end + +lemma doubling_Gamma_eq_Gamma {s : ℝ} (hs : 0 < s) : doubling_Gamma s = Gamma s := +begin + refine eq_Gamma_of_log_convex doubling_Gamma_log_convex_Ioi + (λ y hy, doubling_Gamma_add_one y hy.ne') (λ y hy, _) doubling_Gamma_one hs, + apply_rules [mul_pos, Gamma_pos_of_pos, add_pos, inv_pos_of_pos, + rpow_pos_of_pos, two_pos, one_pos, sqrt_pos_of_pos pi_pos] +end + +/-- Legendre's doubling formula for the Gamma function, for positive real arguments. Note that +we shall later prove this for all `s` as `real.Gamma_mul_Gamma_add_half` (superseding this result) +but this result is needed as an intermediate step. -/ +lemma Gamma_mul_Gamma_add_half_of_pos {s : ℝ} (hs : 0 < s) : + Gamma s * Gamma (s + 1 / 2) = Gamma (2 * s) * 2 ^ (1 - 2 * s) * sqrt π := +begin + rw [←(doubling_Gamma_eq_Gamma (mul_pos two_pos hs)), + doubling_Gamma, mul_div_cancel_left _ (two_ne_zero' ℝ), + (by abel : 1 - 2 * s = -(2 * s - 1)), rpow_neg zero_le_two], + field_simp [(sqrt_pos_of_pos pi_pos).ne', (rpow_pos_of_pos two_pos (2 * s - 1)).ne'], + ring, +end + +end doubling + end real From 36938f775671ff28bea1c0310f1608e4afbb22e0 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Tue, 6 Jun 2023 05:16:06 +0000 Subject: [PATCH 1148/1291] chore(*): add mathlib4 synchronization comments (#19159) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.category.Module.monoidal.symmetric` * `algebra.lie.base_change` * `algebra.lie.character` * `algebra.lie.direct_sum` * `algebra.lie.semisimple` * `algebra.lie.skew_adjoint` * `algebra.lie.solvable` * `algebra.module.dedekind_domain` * `algebra.symmetrized` * `analysis.ODE.gronwall` * `analysis.box_integral.partition.measure` * `analysis.calculus.fderiv_analytic` * `analysis.calculus.series` * `analysis.inner_product_space.calculus` * `analysis.inner_product_space.euclidean_dist` * `analysis.special_functions.arsinh` * `analysis.special_functions.bernstein` * `analysis.special_functions.compare_exp` * `analysis.special_functions.complex.log_deriv` * `analysis.special_functions.log.deriv` * `analysis.special_functions.trigonometric.deriv` * `analysis.special_functions.trigonometric.inverse_deriv` * `category_theory.bicategory.natural_transformation` * `category_theory.monoidal.of_chosen_finite_products.symmetric` * `category_theory.monoidal.rigid.functor_category` * `category_theory.monoidal.rigid.of_equivalence` * `category_theory.monoidal.types.symmetric` * `category_theory.sites.canonical` * `data.complex.exponential_bounds` * `data.nat.sqrt_norm_num` * `field_theory.laurent` * `geometry.manifold.instances.real` * `geometry.manifold.instances.units_of_normed_algebra` * `geometry.manifold.metrizable` * `geometry.manifold.smooth_manifold_with_corners` * `group_theory.transfer` * `linear_algebra.quadratic_form.basic` * `measure_theory.integral.bochner` * `measure_theory.integral.vitali_caratheodory` * `measure_theory.measure.haar.normed_space` * `measure_theory.measure.lebesgue.eq_haar` * `number_theory.liouville.measure` * `ring_theory.dedekind_domain.ideal` * `ring_theory.polynomial.bernstein` * `ring_theory.witt_vector.defs` * `ring_theory.witt_vector.structure_polynomial` * `topology.continuous_function.stone_weierstrass` * `topology.continuous_function.weierstrass` --- src/algebra/category/Module/monoidal/symmetric.lean | 3 +++ src/algebra/lie/base_change.lean | 3 +++ src/algebra/lie/character.lean | 3 +++ src/algebra/lie/direct_sum.lean | 3 +++ src/algebra/lie/semisimple.lean | 3 +++ src/algebra/lie/skew_adjoint.lean | 3 +++ src/algebra/lie/solvable.lean | 3 +++ src/algebra/module/dedekind_domain.lean | 3 +++ src/algebra/symmetrized.lean | 3 +++ src/analysis/ODE/gronwall.lean | 3 +++ src/analysis/box_integral/partition/measure.lean | 3 +++ src/analysis/calculus/fderiv_analytic.lean | 3 +++ src/analysis/calculus/series.lean | 3 +++ src/analysis/inner_product_space/calculus.lean | 3 +++ src/analysis/inner_product_space/euclidean_dist.lean | 3 +++ src/analysis/special_functions/arsinh.lean | 3 +++ src/analysis/special_functions/bernstein.lean | 3 +++ src/analysis/special_functions/compare_exp.lean | 3 +++ src/analysis/special_functions/complex/log_deriv.lean | 3 +++ src/analysis/special_functions/log/deriv.lean | 3 +++ src/analysis/special_functions/trigonometric/deriv.lean | 3 +++ .../special_functions/trigonometric/inverse_deriv.lean | 3 +++ src/category_theory/bicategory/natural_transformation.lean | 3 +++ .../monoidal/of_chosen_finite_products/symmetric.lean | 3 +++ src/category_theory/monoidal/rigid/functor_category.lean | 3 +++ src/category_theory/monoidal/rigid/of_equivalence.lean | 3 +++ src/category_theory/monoidal/types/symmetric.lean | 3 +++ src/category_theory/sites/canonical.lean | 3 +++ src/data/complex/exponential_bounds.lean | 3 +++ src/data/nat/sqrt_norm_num.lean | 3 +++ src/field_theory/laurent.lean | 3 +++ src/geometry/manifold/instances/real.lean | 3 +++ src/geometry/manifold/instances/units_of_normed_algebra.lean | 3 +++ src/geometry/manifold/metrizable.lean | 3 +++ src/geometry/manifold/smooth_manifold_with_corners.lean | 3 +++ src/group_theory/transfer.lean | 3 +++ src/linear_algebra/quadratic_form/basic.lean | 3 +++ src/measure_theory/integral/bochner.lean | 3 +++ src/measure_theory/integral/vitali_caratheodory.lean | 3 +++ src/measure_theory/measure/haar/normed_space.lean | 3 +++ src/measure_theory/measure/lebesgue/eq_haar.lean | 3 +++ src/number_theory/liouville/measure.lean | 3 +++ src/ring_theory/dedekind_domain/ideal.lean | 3 +++ src/ring_theory/polynomial/bernstein.lean | 3 +++ src/ring_theory/witt_vector/defs.lean | 3 +++ src/ring_theory/witt_vector/structure_polynomial.lean | 3 +++ src/topology/continuous_function/stone_weierstrass.lean | 3 +++ src/topology/continuous_function/weierstrass.lean | 3 +++ 48 files changed, 144 insertions(+) diff --git a/src/algebra/category/Module/monoidal/symmetric.lean b/src/algebra/category/Module/monoidal/symmetric.lean index 9bf57b263f96c..177cae80fa58d 100644 --- a/src/algebra/category/Module/monoidal/symmetric.lean +++ b/src/algebra/category/Module/monoidal/symmetric.lean @@ -8,6 +8,9 @@ import algebra.category.Module.monoidal.basic /-! # The symmetric monoidal structure on `Module R`. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ universes v w x u diff --git a/src/algebra/lie/base_change.lean b/src/algebra/lie/base_change.lean index 11c0631716dbd..bf0b6beaa9c74 100644 --- a/src/algebra/lie/base_change.lean +++ b/src/algebra/lie/base_change.lean @@ -9,6 +9,9 @@ import algebra.lie.tensor_product /-! # Extension and restriction of scalars for Lie algebras +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Lie algebras have a well-behaved theory of extension and restriction of scalars. ## Main definitions diff --git a/src/algebra/lie/character.lean b/src/algebra/lie/character.lean index 170178ca0b8fe..1d85a4465fa2a 100644 --- a/src/algebra/lie/character.lean +++ b/src/algebra/lie/character.lean @@ -10,6 +10,9 @@ import linear_algebra.dual /-! # Characters of Lie algebras +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A character of a Lie algebra `L` over a commutative ring `R` is a morphism of Lie algebras `L → R`, where `R` is regarded as a Lie algebra over itself via the ring commutator. For an Abelian Lie algebra (e.g., a Cartan subalgebra of a semisimple Lie algebra) a character is just a linear form. diff --git a/src/algebra/lie/direct_sum.lean b/src/algebra/lie/direct_sum.lean index 1b8a6dd8e8ed7..fd1db3af14d08 100644 --- a/src/algebra/lie/direct_sum.lean +++ b/src/algebra/lie/direct_sum.lean @@ -11,6 +11,9 @@ import algebra.lie.basic /-! # Direct sums of Lie algebras and Lie modules +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Direct sums of Lie algebras and Lie modules carry natural algebra and module structures. ## Tags diff --git a/src/algebra/lie/semisimple.lean b/src/algebra/lie/semisimple.lean index 60d72815ec0ab..d76852391a9c7 100644 --- a/src/algebra/lie/semisimple.lean +++ b/src/algebra/lie/semisimple.lean @@ -8,6 +8,9 @@ import algebra.lie.solvable /-! # Semisimple Lie algebras +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The famous Cartan-Dynkin-Killing classification of semisimple Lie algebras renders them one of the most important classes of Lie algebras. In this file we define simple and semisimple Lie algebras and prove some basic related results. diff --git a/src/algebra/lie/skew_adjoint.lean b/src/algebra/lie/skew_adjoint.lean index 4242eeeb37ec8..c9f78e816a6d0 100644 --- a/src/algebra/lie/skew_adjoint.lean +++ b/src/algebra/lie/skew_adjoint.lean @@ -9,6 +9,9 @@ import linear_algebra.matrix.bilinear_form /-! # Lie algebras of skew-adjoint endomorphisms of a bilinear form +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + When a module carries a bilinear form, the Lie algebra of endomorphisms of the module contains a distinguished Lie subalgebra: the skew-adjoint endomorphisms. Such subalgebras are important because they provide a simple, explicit construction of the so-called classical Lie algebras. diff --git a/src/algebra/lie/solvable.lean b/src/algebra/lie/solvable.lean index 463baa24be718..e463807199cc5 100644 --- a/src/algebra/lie/solvable.lean +++ b/src/algebra/lie/solvable.lean @@ -10,6 +10,9 @@ import order.hom.basic /-! # Solvable Lie algebras +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Like groups, Lie algebras admit a natural concept of solvability. We define this here via the derived series and prove some related results. We also define the radical of a Lie algebra and prove that it is solvable when the Lie algebra is Noetherian. diff --git a/src/algebra/module/dedekind_domain.lean b/src/algebra/module/dedekind_domain.lean index eee85d418a999..d766ff9bc44da 100644 --- a/src/algebra/module/dedekind_domain.lean +++ b/src/algebra/module/dedekind_domain.lean @@ -9,6 +9,9 @@ import ring_theory.dedekind_domain.ideal /-! # Modules over a Dedekind domain +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Over a Dedekind domain, a `I`-torsion module is the internal direct sum of its `p i ^ e i`-torsion submodules, where `I = ∏ i, p i ^ e i` is its unique decomposition in prime ideals. Therefore, as any finitely generated torsion module is `I`-torsion for some `I`, it is an internal diff --git a/src/algebra/symmetrized.lean b/src/algebra/symmetrized.lean index c335e5b9864c5..f4e94a886498f 100644 --- a/src/algebra/symmetrized.lean +++ b/src/algebra/symmetrized.lean @@ -9,6 +9,9 @@ import algebra.module.basic /-! # Symmetrized algebra +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A commutative multiplication on a real or complex space can be constructed from any multiplication by "symmetrization" i.e $$ diff --git a/src/analysis/ODE/gronwall.lean b/src/analysis/ODE/gronwall.lean index 04426303cd2c0..3c9bd852517ed 100644 --- a/src/analysis/ODE/gronwall.lean +++ b/src/analysis/ODE/gronwall.lean @@ -8,6 +8,9 @@ import analysis.special_functions.exp_deriv /-! # Grönwall's inequality +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The main technical result of this file is the Grönwall-like inequality `norm_le_gronwall_bound_of_norm_deriv_right_le`. It states that if `f : ℝ → E` satisfies `‖f a‖ ≤ δ` and `∀ x ∈ [a, b), ‖f' x‖ ≤ K * ‖f x‖ + ε`, then for all `x ∈ [a, b]` we have `‖f x‖ ≤ δ * exp (K * diff --git a/src/analysis/box_integral/partition/measure.lean b/src/analysis/box_integral/partition/measure.lean index e61daf8b963c1..ad540c3aab4ce 100644 --- a/src/analysis/box_integral/partition/measure.lean +++ b/src/analysis/box_integral/partition/measure.lean @@ -9,6 +9,9 @@ import measure_theory.measure.lebesgue.basic /-! # Box-additive functions defined by measures +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove a few simple facts about rectangular boxes, partitions, and measures: - given a box `I : box ι`, its coercion to `set (ι → ℝ)` and `I.Icc` are measurable sets; diff --git a/src/analysis/calculus/fderiv_analytic.lean b/src/analysis/calculus/fderiv_analytic.lean index 9b2775354f2fe..59431d132384c 100644 --- a/src/analysis/calculus/fderiv_analytic.lean +++ b/src/analysis/calculus/fderiv_analytic.lean @@ -10,6 +10,9 @@ import analysis.calculus.cont_diff_def /-! # Frechet derivatives of analytic functions. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A function expressible as a power series at a point has a Frechet derivative there. Also the special case in terms of `deriv` when the domain is 1-dimensional. -/ diff --git a/src/analysis/calculus/series.lean b/src/analysis/calculus/series.lean index a5028c8f9f57b..cc4856796a27a 100644 --- a/src/analysis/calculus/series.lean +++ b/src/analysis/calculus/series.lean @@ -10,6 +10,9 @@ import data.nat.cast.with_top /-! # Smoothness of series +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We show that series of functions are continuous, or differentiable, or smooth, when each individual function in the series is and additionally suitable uniform summable bounds are satisfied. diff --git a/src/analysis/inner_product_space/calculus.lean b/src/analysis/inner_product_space/calculus.lean index 31ddee695201e..74af70e275aee 100644 --- a/src/analysis/inner_product_space/calculus.lean +++ b/src/analysis/inner_product_space/calculus.lean @@ -9,6 +9,9 @@ import analysis.special_functions.sqrt /-! # Calculus in inner product spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove that the inner product and square of the norm in an inner space are infinitely `ℝ`-smooth. In order to state these results, we need a `normed_space ℝ E` instance. Though we can deduce this structure from `inner_product_space 𝕜 E`, this instance may be diff --git a/src/analysis/inner_product_space/euclidean_dist.lean b/src/analysis/inner_product_space/euclidean_dist.lean index d8a0b8fb79283..7cc795527ac2c 100644 --- a/src/analysis/inner_product_space/euclidean_dist.lean +++ b/src/analysis/inner_product_space/euclidean_dist.lean @@ -9,6 +9,9 @@ import analysis.inner_product_space.pi_L2 /-! # Euclidean distance on a finite dimensional space +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + When we define a smooth bump function on a normed space, it is useful to have a smooth distance on the space. Since the default distance is not guaranteed to be smooth, we define `to_euclidean` to be an equivalence between a finite dimensional topological vector space and the standard Euclidean diff --git a/src/analysis/special_functions/arsinh.lean b/src/analysis/special_functions/arsinh.lean index 37231a241ee6d..9611925d171fd 100644 --- a/src/analysis/special_functions/arsinh.lean +++ b/src/analysis/special_functions/arsinh.lean @@ -9,6 +9,9 @@ import analysis.special_functions.log.basic /-! # Inverse of the sinh function +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove that sinh is bijective and hence has an inverse, arsinh. diff --git a/src/analysis/special_functions/bernstein.lean b/src/analysis/special_functions/bernstein.lean index e26be5fe94937..22e78e0f6d775 100644 --- a/src/analysis/special_functions/bernstein.lean +++ b/src/analysis/special_functions/bernstein.lean @@ -11,6 +11,9 @@ import topology.continuous_function.compact /-! # Bernstein approximations and Weierstrass' theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We prove that the Bernstein approximations ``` ∑ k : fin (n+1), f (k/n : ℝ) * n.choose k * x^k * (1-x)^(n-k) diff --git a/src/analysis/special_functions/compare_exp.lean b/src/analysis/special_functions/compare_exp.lean index 87fb46b1c9c32..c48a5ca404fbd 100644 --- a/src/analysis/special_functions/compare_exp.lean +++ b/src/analysis/special_functions/compare_exp.lean @@ -10,6 +10,9 @@ import analysis.asymptotics.specific_asymptotics /-! # Growth estimates on `x ^ y` for complex `x`, `y` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Let `l` be a filter on `ℂ` such that `complex.re` tends to infinity along `l` and `complex.im z` grows at a subexponential rate compared to `complex.re z`. Then diff --git a/src/analysis/special_functions/complex/log_deriv.lean b/src/analysis/special_functions/complex/log_deriv.lean index 139d31bbdbe4a..ce07999d4f09f 100644 --- a/src/analysis/special_functions/complex/log_deriv.lean +++ b/src/analysis/special_functions/complex/log_deriv.lean @@ -10,6 +10,9 @@ import analysis.special_functions.exp_deriv /-! # Differentiability of the complex `log` function +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + -/ noncomputable theory diff --git a/src/analysis/special_functions/log/deriv.lean b/src/analysis/special_functions/log/deriv.lean index cc1b00d9942ed..6f006ca1234f3 100644 --- a/src/analysis/special_functions/log/deriv.lean +++ b/src/analysis/special_functions/log/deriv.lean @@ -11,6 +11,9 @@ import analysis.special_functions.exp_deriv /-! # Derivative and series expansion of real logarithm +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove that `real.log` is infinitely smooth at all nonzero `x : ℝ`. We also prove that the series `∑' n : ℕ, x ^ (n + 1) / (n + 1)` converges to `(-real.log (1 - x))` for all `x : ℝ`, `|x| < 1`. diff --git a/src/analysis/special_functions/trigonometric/deriv.lean b/src/analysis/special_functions/trigonometric/deriv.lean index 29a599126d297..65fdd446450b5 100644 --- a/src/analysis/special_functions/trigonometric/deriv.lean +++ b/src/analysis/special_functions/trigonometric/deriv.lean @@ -10,6 +10,9 @@ import analysis.special_functions.trigonometric.basic /-! # Differentiability of trigonometric functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main statements The differentiability of the usual trigonometric functions is proved, and their derivatives are diff --git a/src/analysis/special_functions/trigonometric/inverse_deriv.lean b/src/analysis/special_functions/trigonometric/inverse_deriv.lean index bdff998aeb4d3..ac904bb69220b 100644 --- a/src/analysis/special_functions/trigonometric/inverse_deriv.lean +++ b/src/analysis/special_functions/trigonometric/inverse_deriv.lean @@ -9,6 +9,9 @@ import analysis.special_functions.trigonometric.deriv /-! # derivatives of the inverse trigonometric functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Derivatives of `arcsin` and `arccos`. -/ diff --git a/src/category_theory/bicategory/natural_transformation.lean b/src/category_theory/bicategory/natural_transformation.lean index 08f8db6d452fe..79d7e6720944e 100644 --- a/src/category_theory/bicategory/natural_transformation.lean +++ b/src/category_theory/bicategory/natural_transformation.lean @@ -8,6 +8,9 @@ import category_theory.bicategory.functor /-! # Oplax natural transformations +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Just as there are natural transformations between functors, there are oplax natural transformations between oplax functors. The equality in the naturality of natural transformations is replaced by a specified 2-morphism `F.map f ≫ app b ⟶ app a ≫ G.map f` in the case of oplax natural diff --git a/src/category_theory/monoidal/of_chosen_finite_products/symmetric.lean b/src/category_theory/monoidal/of_chosen_finite_products/symmetric.lean index 7200c53a49fd8..a7e27822045b5 100644 --- a/src/category_theory/monoidal/of_chosen_finite_products/symmetric.lean +++ b/src/category_theory/monoidal/of_chosen_finite_products/symmetric.lean @@ -9,6 +9,9 @@ import category_theory.monoidal.of_chosen_finite_products.basic /-! # The symmetric monoidal structure on a category with chosen finite products. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + -/ universes v u diff --git a/src/category_theory/monoidal/rigid/functor_category.lean b/src/category_theory/monoidal/rigid/functor_category.lean index ef3315f137b3d..61e3a4d4a5767 100644 --- a/src/category_theory/monoidal/rigid/functor_category.lean +++ b/src/category_theory/monoidal/rigid/functor_category.lean @@ -9,6 +9,9 @@ import category_theory.monoidal.functor_category /-! # Functors from a groupoid into a right/left rigid category form a right/left rigid category. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + (Using the pointwise monoidal structure on the functor category.) -/ diff --git a/src/category_theory/monoidal/rigid/of_equivalence.lean b/src/category_theory/monoidal/rigid/of_equivalence.lean index 567fa184962d7..6646a63db923c 100644 --- a/src/category_theory/monoidal/rigid/of_equivalence.lean +++ b/src/category_theory/monoidal/rigid/of_equivalence.lean @@ -7,6 +7,9 @@ import category_theory.monoidal.rigid.basic /-! # Transport rigid structures over a monoidal equivalence. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ noncomputable theory diff --git a/src/category_theory/monoidal/types/symmetric.lean b/src/category_theory/monoidal/types/symmetric.lean index 3aab388e36ba6..c8d88e799b7e7 100644 --- a/src/category_theory/monoidal/types/symmetric.lean +++ b/src/category_theory/monoidal/types/symmetric.lean @@ -8,6 +8,9 @@ import category_theory.monoidal.types.basic /-! # The category of types is a symmetric monoidal category + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ open category_theory diff --git a/src/category_theory/sites/canonical.lean b/src/category_theory/sites/canonical.lean index cd8ea2aa40f57..f92fdc076e0f1 100644 --- a/src/category_theory/sites/canonical.lean +++ b/src/category_theory/sites/canonical.lean @@ -10,6 +10,9 @@ import category_theory.sites.sheaf_of_types /-! # The canonical topology on a category +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define the finest (largest) Grothendieck topology for which a given presheaf `P` is a sheaf. This is well defined since if `P` is a sheaf for a topology `J`, then it is a sheaf for any coarser (smaller) topology. Nonetheless we define the topology explicitly by specifying its sieves: diff --git a/src/data/complex/exponential_bounds.lean b/src/data/complex/exponential_bounds.lean index 76cbd13584f41..53d37bc86f1be 100644 --- a/src/data/complex/exponential_bounds.lean +++ b/src/data/complex/exponential_bounds.lean @@ -9,6 +9,9 @@ import analysis.special_functions.log.deriv /-! # Bounds on specific values of the exponential + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ namespace real diff --git a/src/data/nat/sqrt_norm_num.lean b/src/data/nat/sqrt_norm_num.lean index 8f63384879b02..4ae37dded19f0 100644 --- a/src/data/nat/sqrt_norm_num.lean +++ b/src/data/nat/sqrt_norm_num.lean @@ -8,6 +8,9 @@ import data.nat.sqrt /-! ### `norm_num` plugin for `sqrt` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The `norm_num` plugin evaluates `sqrt` by bounding it between consecutive integers. -/ diff --git a/src/field_theory/laurent.lean b/src/field_theory/laurent.lean index d504346dd9d9f..2e39eea2aeab5 100644 --- a/src/field_theory/laurent.lean +++ b/src/field_theory/laurent.lean @@ -10,6 +10,9 @@ import field_theory.ratfunc /-! # Laurent expansions of rational functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main declarations * `ratfunc.laurent`: the Laurent expansion of the rational function `f` at `r`, as an `alg_hom`. diff --git a/src/geometry/manifold/instances/real.lean b/src/geometry/manifold/instances/real.lean index 533dcf5c7af9a..f82c76fdf1478 100644 --- a/src/geometry/manifold/instances/real.lean +++ b/src/geometry/manifold/instances/real.lean @@ -9,6 +9,9 @@ import analysis.inner_product_space.pi_L2 /-! # Constructing examples of manifolds over ℝ +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We introduce the necessary bits to be able to define manifolds modelled over `ℝ^n`, boundaryless or with boundary or with corners. As a concrete example, we construct explicitly the manifold with boundary structure on the real interval `[x, y]`. diff --git a/src/geometry/manifold/instances/units_of_normed_algebra.lean b/src/geometry/manifold/instances/units_of_normed_algebra.lean index 77bfb586b575e..a7b7b50267a71 100644 --- a/src/geometry/manifold/instances/units_of_normed_algebra.lean +++ b/src/geometry/manifold/instances/units_of_normed_algebra.lean @@ -10,6 +10,9 @@ import analysis.normed_space.units /-! # Units of a normed algebra +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file is a stub, containing a construction of the charted space structure on the group of units of a complete normed ring `R`, and of the smooth manifold structure on the group of units of a complete normed `𝕜`-algebra `R`. diff --git a/src/geometry/manifold/metrizable.lean b/src/geometry/manifold/metrizable.lean index 432c3b5b1dae0..681a32a20b59e 100644 --- a/src/geometry/manifold/metrizable.lean +++ b/src/geometry/manifold/metrizable.lean @@ -10,6 +10,9 @@ import topology.metric_space.metrizable /-! # Metrizability of a σ-compact manifold +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we show that a σ-compact Hausdorff topological manifold over a finite dimensional real vector space is metrizable. -/ diff --git a/src/geometry/manifold/smooth_manifold_with_corners.lean b/src/geometry/manifold/smooth_manifold_with_corners.lean index ce60df1b49399..9eade2cf3a81a 100644 --- a/src/geometry/manifold/smooth_manifold_with_corners.lean +++ b/src/geometry/manifold/smooth_manifold_with_corners.lean @@ -9,6 +9,9 @@ import geometry.manifold.charted_space /-! # Smooth manifolds (possibly with boundary or corners) +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A smooth manifold is a manifold modelled on a normed vector space, or a subset like a half-space (to get manifolds with boundaries) for which the changes of coordinates are smooth maps. We define a model with corners as a map `I : H → E` embedding nicely the topological space `H` in diff --git a/src/group_theory/transfer.lean b/src/group_theory/transfer.lean index 3bc1bc754c47a..f54f4a3cc4f1d 100644 --- a/src/group_theory/transfer.lean +++ b/src/group_theory/transfer.lean @@ -10,6 +10,9 @@ import group_theory.sylow /-! # The Transfer Homomorphism +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we construct the transfer homomorphism. ## Main definitions diff --git a/src/linear_algebra/quadratic_form/basic.lean b/src/linear_algebra/quadratic_form/basic.lean index 61d4c6bda1277..cea7b02947391 100644 --- a/src/linear_algebra/quadratic_form/basic.lean +++ b/src/linear_algebra/quadratic_form/basic.lean @@ -12,6 +12,9 @@ import linear_algebra.matrix.symmetric /-! # Quadratic forms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines quadratic forms over a `R`-module `M`. A quadratic form on a ring `R` is a map `Q : M → R` such that: * `quadratic_form.map_smul`: `Q (a • x) = a * a * Q x` diff --git a/src/measure_theory/integral/bochner.lean b/src/measure_theory/integral/bochner.lean index e2627ba77641b..c594e1e066aed 100644 --- a/src/measure_theory/integral/bochner.lean +++ b/src/measure_theory/integral/bochner.lean @@ -8,6 +8,9 @@ import measure_theory.integral.set_to_l1 /-! # Bochner integral +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The Bochner integral extends the definition of the Lebesgue integral to functions that map from a measure space into a Banach space (complete normed vector space). It is constructed here by extending the integral on simple functions. diff --git a/src/measure_theory/integral/vitali_caratheodory.lean b/src/measure_theory/integral/vitali_caratheodory.lean index e5e59cfb7fbdd..2de96efc96d3d 100644 --- a/src/measure_theory/integral/vitali_caratheodory.lean +++ b/src/measure_theory/integral/vitali_caratheodory.lean @@ -12,6 +12,9 @@ import topology.instances.ereal /-! # Vitali-Carathéodory theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Vitali-Carathéodory theorem asserts the following. Consider an integrable function `f : α → ℝ` on a space with a regular measure. Then there exists a function `g : α → ereal` such that `f x < g x` everywhere, `g` is lower semicontinuous, and the integral of `g` is arbitrarily close to that of diff --git a/src/measure_theory/measure/haar/normed_space.lean b/src/measure_theory/measure/haar/normed_space.lean index 49aae295bb8dd..f4cae46185949 100644 --- a/src/measure_theory/measure/haar/normed_space.lean +++ b/src/measure_theory/measure/haar/normed_space.lean @@ -9,6 +9,9 @@ import measure_theory.integral.bochner /-! # Basic properties of Haar measures on real vector spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + -/ noncomputable theory diff --git a/src/measure_theory/measure/lebesgue/eq_haar.lean b/src/measure_theory/measure/lebesgue/eq_haar.lean index 1bdac70427224..649bd355f7b6a 100644 --- a/src/measure_theory/measure/lebesgue/eq_haar.lean +++ b/src/measure_theory/measure/lebesgue/eq_haar.lean @@ -13,6 +13,9 @@ import measure_theory.measure.doubling /-! # Relationship between the Haar and Lebesgue measures +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We prove that the Haar measure and Lebesgue measure are equal on `ℝ` and on `ℝ^ι`, in `measure_theory.add_haar_measure_eq_volume` and `measure_theory.add_haar_measure_eq_volume_pi`. diff --git a/src/number_theory/liouville/measure.lean b/src/number_theory/liouville/measure.lean index 8e5eb260b4453..b8fc7040fa943 100644 --- a/src/number_theory/liouville/measure.lean +++ b/src/number_theory/liouville/measure.lean @@ -11,6 +11,9 @@ import analysis.p_series /-! # Volume of the set of Liouville numbers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove that the set of Liouville numbers with exponent (irrationality measure) strictly greater than two is a set of Lebesuge measure zero, see `volume_Union_set_of_liouville_with`. diff --git a/src/ring_theory/dedekind_domain/ideal.lean b/src/ring_theory/dedekind_domain/ideal.lean index 53da5873c2207..be1aed2e252fd 100644 --- a/src/ring_theory/dedekind_domain/ideal.lean +++ b/src/ring_theory/dedekind_domain/ideal.lean @@ -15,6 +15,9 @@ import ring_theory.chain_of_divisors /-! # Dedekind domains and ideals +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we show a ring is a Dedekind domain iff all fractional ideals are invertible. Then we prove some results on the unique factorization monoid structure of the ideals. diff --git a/src/ring_theory/polynomial/bernstein.lean b/src/ring_theory/polynomial/bernstein.lean index cd545d889dc7f..bbf4b6078196e 100644 --- a/src/ring_theory/polynomial/bernstein.lean +++ b/src/ring_theory/polynomial/bernstein.lean @@ -13,6 +13,9 @@ import data.mv_polynomial.pderiv /-! # Bernstein polynomials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The definition of the Bernstein polynomials ``` bernstein_polynomial (R : Type*) [comm_ring R] (n ν : ℕ) : R[X] := diff --git a/src/ring_theory/witt_vector/defs.lean b/src/ring_theory/witt_vector/defs.lean index 07648860fc301..c3ba171fd8dfa 100644 --- a/src/ring_theory/witt_vector/defs.lean +++ b/src/ring_theory/witt_vector/defs.lean @@ -9,6 +9,9 @@ import ring_theory.witt_vector.structure_polynomial /-! # Witt vectors +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define the type of `p`-typical Witt vectors and ring operations on it. The ring axioms are verified in `ring_theory/witt_vector/basic.lean`. diff --git a/src/ring_theory/witt_vector/structure_polynomial.lean b/src/ring_theory/witt_vector/structure_polynomial.lean index 522b51afd6e7f..0362bd35f4a7f 100644 --- a/src/ring_theory/witt_vector/structure_polynomial.lean +++ b/src/ring_theory/witt_vector/structure_polynomial.lean @@ -11,6 +11,9 @@ import ring_theory.witt_vector.witt_polynomial /-! # Witt structure polynomials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove the main theorem that makes the whole theory of Witt vectors work. Briefly, consider a polynomial `Φ : mv_polynomial idx ℤ` over the integers, with polynomials variables indexed by an arbitrary type `idx`. diff --git a/src/topology/continuous_function/stone_weierstrass.lean b/src/topology/continuous_function/stone_weierstrass.lean index fe7cdff481cbe..4d8ecabec58db 100644 --- a/src/topology/continuous_function/stone_weierstrass.lean +++ b/src/topology/continuous_function/stone_weierstrass.lean @@ -9,6 +9,9 @@ import data.is_R_or_C.basic /-! # The Stone-Weierstrass theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + If a subalgebra `A` of `C(X, ℝ)`, where `X` is a compact topological space, separates points, then it is dense. diff --git a/src/topology/continuous_function/weierstrass.lean b/src/topology/continuous_function/weierstrass.lean index d60d6f9f63f88..0f90c5b6c3135 100644 --- a/src/topology/continuous_function/weierstrass.lean +++ b/src/topology/continuous_function/weierstrass.lean @@ -9,6 +9,9 @@ import topology.algebra.algebra /-! # The Weierstrass approximation theorem for continuous functions on `[a,b]` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We've already proved the Weierstrass approximation theorem in the sense that we've shown that the Bernstein approximations to a continuous function on `[0,1]` converge uniformly. From 58a272265b5e05f258161260dd2c5d247213cbd3 Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Tue, 6 Jun 2023 17:18:51 +0000 Subject: [PATCH 1149/1291] chore(algebra/algebra/spectrum): split file (#19161) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This splits off all the material related to polynomials from `algebra.algebra.spectrum` into a new file `field_theory.is_alg_closed.spectrum`, because (almost) all of it requires `is_alg_closed 𝕜`. This significantly simplifies the import tree for this file, and for a few files which import it. This also does two minor housekeeping chores: 1. generalizes type class assumptions for `alg_hom.mem_resolvent_set_apply` and `alg_hom.spectrum_apply_subset` 2. rephrases `spectrum.nonempty_of_is_alg_closed_of_finite_dimensional` to use `set.nonempty`. --- src/algebra/algebra/spectrum.lean | 110 +------------ src/analysis/normed_space/spectrum.lean | 2 +- src/category_theory/preadditive/schur.lean | 2 +- src/field_theory/is_alg_closed/spectrum.lean | 163 +++++++++++++++++++ src/linear_algebra/eigenspace.lean | 2 +- 5 files changed, 167 insertions(+), 112 deletions(-) create mode 100644 src/field_theory/is_alg_closed/spectrum.lean diff --git a/src/algebra/algebra/spectrum.lean b/src/algebra/algebra/spectrum.lean index 457dbacc3e6a6..cdbf09f8e3bd4 100644 --- a/src/algebra/algebra/spectrum.lean +++ b/src/algebra/algebra/spectrum.lean @@ -5,7 +5,6 @@ Authors: Jireh Loreaux -/ import algebra.star.pointwise import algebra.star.subalgebra -import field_theory.is_alg_closed.basic import tactic.noncomm_ring /-! # Spectrum of an element in an algebra @@ -30,8 +29,6 @@ This theory will serve as the foundation for spectral theory in Banach algebras. units (of `R`) in `σ (a*b)` coincide with those in `σ (b*a)`. * `spectrum.scalar_eq`: in a nontrivial algebra over a field, the spectrum of a scalar is a singleton. -* `spectrum.subset_polynomial_aeval`, `spectrum.map_polynomial_aeval_of_degree_pos`, - `spectrum.map_polynomial_aeval_of_nonempty`: variations on the spectral mapping theorem. ## Notations @@ -86,7 +83,6 @@ noncomputable def is_unit.sub_inv_smul {r : Rˣ} {s : R} {a : A} end defs namespace spectrum -open_locale polynomial section scalar_semiring @@ -200,8 +196,6 @@ lemma smul_mem_smul_iff {a : A} {s : R} {r : Rˣ} : by simp only [mem_iff, not_iff_not, algebra.algebra_map_eq_smul_one, smul_assoc, ←smul_sub, is_unit_smul_iff] -open_locale polynomial - theorem unit_smul_eq_smul (a : A) (r : Rˣ) : σ (r • a) = r • σ a := begin @@ -290,20 +284,6 @@ lemma sub_singleton_eq (a : A) (r : R) : (σ a) - {r} = σ (a - ↑ₐr) := by simpa only [neg_sub, neg_eq] using congr_arg has_neg.neg (singleton_sub_eq a r) -open polynomial - -lemma exists_mem_of_not_is_unit_aeval_prod [is_domain R] {p : R[X]} {a : A} (hp : p ≠ 0) - (h : ¬is_unit (aeval a (multiset.map (λ (x : R), X - C x) p.roots).prod)) : - ∃ k : R, k ∈ σ a ∧ eval k p = 0 := -begin - rw [←multiset.prod_to_list, alg_hom.map_list_prod] at h, - replace h := mt list.prod_is_unit h, - simp only [not_forall, exists_prop, aeval_C, multiset.mem_to_list, - list.mem_map, aeval_X, exists_exists_and_eq_and, multiset.mem_map, alg_hom.map_sub] at h, - rcases h with ⟨r, r_mem, r_nu⟩, - exact ⟨r, by rwa [mem_iff, ←is_unit.sub_iff], by rwa [←is_root.def, ←mem_roots hp]⟩ -end - end scalar_ring section scalar_field @@ -364,94 +344,6 @@ begin simpa only [units.coe_inv] using inv_mem_iff.mp hk, } end -open polynomial -/-- Half of the spectral mapping theorem for polynomials. We prove it separately -because it holds over any field, whereas `spectrum.map_polynomial_aeval_of_degree_pos` and -`spectrum.map_polynomial_aeval_of_nonempty` need the field to be algebraically closed. -/ -theorem subset_polynomial_aeval (a : A) (p : 𝕜[X]) : - (λ k, eval k p) '' (σ a) ⊆ σ (aeval a p) := -begin - rintros _ ⟨k, hk, rfl⟩, - let q := C (eval k p) - p, - have hroot : is_root q k, by simp only [eval_C, eval_sub, sub_self, is_root.def], - rw [←mul_div_eq_iff_is_root, ←neg_mul_neg, neg_sub] at hroot, - have aeval_q_eq : ↑ₐ(eval k p) - aeval a p = aeval a q, - by simp only [aeval_C, alg_hom.map_sub, sub_left_inj], - rw [mem_iff, aeval_q_eq, ←hroot, aeval_mul], - have hcomm := (commute.all (C k - X) (- (q / (X - C k)))).map (aeval a), - apply mt (λ h, (hcomm.is_unit_mul_iff.mp h).1), - simpa only [aeval_X, aeval_C, alg_hom.map_sub] using hk, -end - -/-- The *spectral mapping theorem* for polynomials. Note: the assumption `degree p > 0` -is necessary in case `σ a = ∅`, for then the left-hand side is `∅` and the right-hand side, -assuming `[nontrivial A]`, is `{k}` where `p = polynomial.C k`. -/ -theorem map_polynomial_aeval_of_degree_pos [is_alg_closed 𝕜] (a : A) (p : 𝕜[X]) - (hdeg : 0 < degree p) : σ (aeval a p) = (λ k, eval k p) '' (σ a) := -begin - /- handle the easy direction via `spectrum.subset_polynomial_aeval` -/ - refine set.eq_of_subset_of_subset (λ k hk, _) (subset_polynomial_aeval a p), - /- write `C k - p` product of linear factors and a constant; show `C k - p ≠ 0`. -/ - have hprod := eq_prod_roots_of_splits_id (is_alg_closed.splits (C k - p)), - have h_ne : C k - p ≠ 0, from ne_zero_of_degree_gt - (by rwa [degree_sub_eq_right_of_degree_lt (lt_of_le_of_lt degree_C_le hdeg)]), - have lead_ne := leading_coeff_ne_zero.mpr h_ne, - have lead_unit := (units.map (↑ₐ).to_monoid_hom (units.mk0 _ lead_ne)).is_unit, - /- leading coefficient is a unit so product of linear factors is not a unit; - apply `exists_mem_of_not_is_unit_aeval_prod`. -/ - have p_a_eq : aeval a (C k - p) = ↑ₐk - aeval a p, - by simp only [aeval_C, alg_hom.map_sub, sub_left_inj], - rw [mem_iff, ←p_a_eq, hprod, aeval_mul, - ((commute.all _ _).map (aeval a)).is_unit_mul_iff, aeval_C] at hk, - replace hk := exists_mem_of_not_is_unit_aeval_prod h_ne (not_and.mp hk lead_unit), - rcases hk with ⟨r, r_mem, r_ev⟩, - exact ⟨r, r_mem, symm (by simpa [eval_sub, eval_C, sub_eq_zero] using r_ev)⟩, -end - -/-- In this version of the spectral mapping theorem, we assume the spectrum -is nonempty instead of assuming the degree of the polynomial is positive. -/ -theorem map_polynomial_aeval_of_nonempty [is_alg_closed 𝕜] (a : A) (p : 𝕜[X]) - (hnon : (σ a).nonempty) : σ (aeval a p) = (λ k, eval k p) '' (σ a) := -begin - nontriviality A, - refine or.elim (le_or_gt (degree p) 0) (λ h, _) (map_polynomial_aeval_of_degree_pos a p), - { rw eq_C_of_degree_le_zero h, - simp only [set.image_congr, eval_C, aeval_C, scalar_eq, set.nonempty.image_const hnon] }, -end - -/-- A specialization of `spectrum.subset_polynomial_aeval` to monic monomials for convenience. -/ -lemma pow_image_subset (a : A) (n : ℕ) : (λ x, x ^ n) '' (σ a) ⊆ σ (a ^ n) := -by simpa only [eval_pow, eval_X, aeval_X_pow] using subset_polynomial_aeval a (X ^ n : 𝕜[X]) - -/-- A specialization of `spectrum.map_polynomial_aeval_of_nonempty` to monic monomials for -convenience. -/ -lemma map_pow_of_pos [is_alg_closed 𝕜] (a : A) {n : ℕ} (hn : 0 < n) : - σ (a ^ n) = (λ x, x ^ n) '' (σ a) := -by simpa only [aeval_X_pow, eval_pow, eval_X] using - map_polynomial_aeval_of_degree_pos a (X ^ n : 𝕜[X]) (by { rw_mod_cast degree_X_pow, exact hn }) - -/-- A specialization of `spectrum.map_polynomial_aeval_of_nonempty` to monic monomials for -convenience. -/ -lemma map_pow_of_nonempty [is_alg_closed 𝕜] {a : A} (ha : (σ a).nonempty) (n : ℕ) : - σ (a ^ n) = (λ x, x ^ n) '' (σ a) := -by simpa only [aeval_X_pow, eval_pow, eval_X] using map_polynomial_aeval_of_nonempty a (X ^ n) ha - -variable (𝕜) -/-- -Every element `a` in a nontrivial finite-dimensional algebra `A` -over an algebraically closed field `𝕜` has non-empty spectrum. -/ --- We will use this both to show eigenvalues exist, and to prove Schur's lemma. -lemma nonempty_of_is_alg_closed_of_finite_dimensional [is_alg_closed 𝕜] - [nontrivial A] [I : finite_dimensional 𝕜 A] (a : A) : - ∃ k : 𝕜, k ∈ σ a := -begin - obtain ⟨p, ⟨h_mon, h_eval_p⟩⟩ := is_integral_of_noetherian (is_noetherian.iff_fg.2 I) a, - have nu : ¬ is_unit (aeval a p), { rw [←aeval_def] at h_eval_p, rw h_eval_p, simp, }, - rw [eq_prod_roots_of_monic_of_splits_id h_mon (is_alg_closed.splits p)] at nu, - obtain ⟨k, hk, _⟩ := exists_mem_of_not_is_unit_aeval_prod (monic.ne_zero h_mon) nu, - exact ⟨k, hk⟩ -end - end scalar_field end spectrum @@ -460,7 +352,7 @@ namespace alg_hom section comm_semiring -variables {F R A B : Type*} [comm_ring R] [ring A] [algebra R A] [ring B] [algebra R B] +variables {F R A B : Type*} [comm_semiring R] [ring A] [algebra R A] [ring B] [algebra R B] variables [alg_hom_class F R A B] local notation `σ` := spectrum R local notation `↑ₐ` := algebra_map R A diff --git a/src/analysis/normed_space/spectrum.lean b/src/analysis/normed_space/spectrum.lean index 13a1e0bf880d1..8fcc2fedf5890 100644 --- a/src/analysis/normed_space/spectrum.lean +++ b/src/analysis/normed_space/spectrum.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Jireh Loreaux. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jireh Loreaux -/ -import algebra.algebra.spectrum +import field_theory.is_alg_closed.spectrum import analysis.complex.liouville import analysis.complex.polynomial import analysis.analytic.radius_liminf diff --git a/src/category_theory/preadditive/schur.lean b/src/category_theory/preadditive/schur.lean index 533e1234c484f..f2a2125b233c6 100644 --- a/src/category_theory/preadditive/schur.lean +++ b/src/category_theory/preadditive/schur.lean @@ -7,7 +7,7 @@ import algebra.group.ext import category_theory.simple import category_theory.linear.basic import category_theory.endomorphism -import algebra.algebra.spectrum +import field_theory.is_alg_closed.spectrum /-! # Schur's lemma diff --git a/src/field_theory/is_alg_closed/spectrum.lean b/src/field_theory/is_alg_closed/spectrum.lean new file mode 100644 index 0000000000000..f63674be31c92 --- /dev/null +++ b/src/field_theory/is_alg_closed/spectrum.lean @@ -0,0 +1,163 @@ +/- +Copyright (c) 2021 Jireh Loreaux. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jireh Loreaux +-/ +import algebra.algebra.spectrum +import field_theory.is_alg_closed.basic + +/-! +# Spectrum mapping theorem + +This file develops proves the spectral mapping theorem for polynomials over algebraically closed +fields. In particular, if `a` is an element of an `𝕜`-algebra `A` where `𝕜` is a field, and +`p : 𝕜[X]` is a polynomial, then the spectrum of `polynomial.aeval a p` contains the image of the +spectrum of `a` under `(λ k, polynomial.eval k p)`. When `𝕜` is algebraically closed, these are in +fact equal (assuming either that the spectrum of `a` is nonempty or the polynomial has positive +degree), which is the **spectral mapping theorem**. + +In addition, this file contains the fact that every element of a finite dimensional nontrivial +algebra over an algebraically closed field has nonempty spectrum. In particular, this is used in +`module.End.exists_eigenvalue` to show that every linear map from a vector space to itself has an +eigenvalue. + +## Main statements + +* `spectrum.subset_polynomial_aeval`, `spectrum.map_polynomial_aeval_of_degree_pos`, + `spectrum.map_polynomial_aeval_of_nonempty`: variations on the **spectral mapping theorem**. +* `spectrum.nonempty_of_is_alg_closed_of_finite_dimensional`: the spectrum is nonempty for any + element of a nontrivial finite dimensional algebra over an algebraically closed field. + +## Notations + +* `σ a` : `spectrum R a` of `a : A` +-/ + +namespace spectrum + +open set polynomial +open_locale pointwise polynomial + +universes u v + +section scalar_ring + +variables {R : Type u} {A : Type v} +variables [comm_ring R] [ring A] [algebra R A] + +local notation `σ` := spectrum R +local notation `↑ₐ` := algebra_map R A + +lemma exists_mem_of_not_is_unit_aeval_prod [is_domain R] {p : R[X]} {a : A} (hp : p ≠ 0) + (h : ¬is_unit (aeval a (multiset.map (λ (x : R), X - C x) p.roots).prod)) : + ∃ k : R, k ∈ σ a ∧ eval k p = 0 := +begin + rw [←multiset.prod_to_list, alg_hom.map_list_prod] at h, + replace h := mt list.prod_is_unit h, + simp only [not_forall, exists_prop, aeval_C, multiset.mem_to_list, + list.mem_map, aeval_X, exists_exists_and_eq_and, multiset.mem_map, alg_hom.map_sub] at h, + rcases h with ⟨r, r_mem, r_nu⟩, + exact ⟨r, by rwa [mem_iff, ←is_unit.sub_iff], by rwa [←is_root.def, ←mem_roots hp]⟩ +end + +end scalar_ring + +section scalar_field + +variables {𝕜 : Type u} {A : Type v} +variables [field 𝕜] [ring A] [algebra 𝕜 A] + +local notation `σ` := spectrum 𝕜 +local notation `↑ₐ` := algebra_map 𝕜 A + +open polynomial +/-- Half of the spectral mapping theorem for polynomials. We prove it separately +because it holds over any field, whereas `spectrum.map_polynomial_aeval_of_degree_pos` and +`spectrum.map_polynomial_aeval_of_nonempty` need the field to be algebraically closed. -/ +theorem subset_polynomial_aeval (a : A) (p : 𝕜[X]) : + (λ k, eval k p) '' (σ a) ⊆ σ (aeval a p) := +begin + rintros _ ⟨k, hk, rfl⟩, + let q := C (eval k p) - p, + have hroot : is_root q k, by simp only [eval_C, eval_sub, sub_self, is_root.def], + rw [←mul_div_eq_iff_is_root, ←neg_mul_neg, neg_sub] at hroot, + have aeval_q_eq : ↑ₐ(eval k p) - aeval a p = aeval a q, + by simp only [aeval_C, alg_hom.map_sub, sub_left_inj], + rw [mem_iff, aeval_q_eq, ←hroot, aeval_mul], + have hcomm := (commute.all (C k - X) (- (q / (X - C k)))).map (aeval a), + apply mt (λ h, (hcomm.is_unit_mul_iff.mp h).1), + simpa only [aeval_X, aeval_C, alg_hom.map_sub] using hk, +end + +/-- The *spectral mapping theorem* for polynomials. Note: the assumption `degree p > 0` +is necessary in case `σ a = ∅`, for then the left-hand side is `∅` and the right-hand side, +assuming `[nontrivial A]`, is `{k}` where `p = polynomial.C k`. -/ +theorem map_polynomial_aeval_of_degree_pos [is_alg_closed 𝕜] (a : A) (p : 𝕜[X]) + (hdeg : 0 < degree p) : σ (aeval a p) = (λ k, eval k p) '' (σ a) := +begin + /- handle the easy direction via `spectrum.subset_polynomial_aeval` -/ + refine set.eq_of_subset_of_subset (λ k hk, _) (subset_polynomial_aeval a p), + /- write `C k - p` product of linear factors and a constant; show `C k - p ≠ 0`. -/ + have hprod := eq_prod_roots_of_splits_id (is_alg_closed.splits (C k - p)), + have h_ne : C k - p ≠ 0, from ne_zero_of_degree_gt + (by rwa [degree_sub_eq_right_of_degree_lt (lt_of_le_of_lt degree_C_le hdeg)]), + have lead_ne := leading_coeff_ne_zero.mpr h_ne, + have lead_unit := (units.map (↑ₐ).to_monoid_hom (units.mk0 _ lead_ne)).is_unit, + /- leading coefficient is a unit so product of linear factors is not a unit; + apply `exists_mem_of_not_is_unit_aeval_prod`. -/ + have p_a_eq : aeval a (C k - p) = ↑ₐk - aeval a p, + by simp only [aeval_C, alg_hom.map_sub, sub_left_inj], + rw [mem_iff, ←p_a_eq, hprod, aeval_mul, + ((commute.all _ _).map (aeval a)).is_unit_mul_iff, aeval_C] at hk, + replace hk := exists_mem_of_not_is_unit_aeval_prod h_ne (not_and.mp hk lead_unit), + rcases hk with ⟨r, r_mem, r_ev⟩, + exact ⟨r, r_mem, symm (by simpa [eval_sub, eval_C, sub_eq_zero] using r_ev)⟩, +end + +/-- In this version of the spectral mapping theorem, we assume the spectrum +is nonempty instead of assuming the degree of the polynomial is positive. -/ +theorem map_polynomial_aeval_of_nonempty [is_alg_closed 𝕜] (a : A) (p : 𝕜[X]) + (hnon : (σ a).nonempty) : σ (aeval a p) = (λ k, eval k p) '' (σ a) := +begin + nontriviality A, + refine or.elim (le_or_gt (degree p) 0) (λ h, _) (map_polynomial_aeval_of_degree_pos a p), + { rw eq_C_of_degree_le_zero h, + simp only [set.image_congr, eval_C, aeval_C, scalar_eq, set.nonempty.image_const hnon] }, +end + +/-- A specialization of `spectrum.subset_polynomial_aeval` to monic monomials for convenience. -/ +lemma pow_image_subset (a : A) (n : ℕ) : (λ x, x ^ n) '' (σ a) ⊆ σ (a ^ n) := +by simpa only [eval_pow, eval_X, aeval_X_pow] using subset_polynomial_aeval a (X ^ n : 𝕜[X]) + +/-- A specialization of `spectrum.map_polynomial_aeval_of_nonempty` to monic monomials for +convenience. -/ +lemma map_pow_of_pos [is_alg_closed 𝕜] (a : A) {n : ℕ} (hn : 0 < n) : + σ (a ^ n) = (λ x, x ^ n) '' (σ a) := +by simpa only [aeval_X_pow, eval_pow, eval_X] using + map_polynomial_aeval_of_degree_pos a (X ^ n : 𝕜[X]) (by { rw_mod_cast degree_X_pow, exact hn }) + +/-- A specialization of `spectrum.map_polynomial_aeval_of_nonempty` to monic monomials for +convenience. -/ +lemma map_pow_of_nonempty [is_alg_closed 𝕜] {a : A} (ha : (σ a).nonempty) (n : ℕ) : + σ (a ^ n) = (λ x, x ^ n) '' (σ a) := +by simpa only [aeval_X_pow, eval_pow, eval_X] using map_polynomial_aeval_of_nonempty a (X ^ n) ha + +variable (𝕜) +/-- +Every element `a` in a nontrivial finite-dimensional algebra `A` +over an algebraically closed field `𝕜` has non-empty spectrum. -/ +-- We will use this both to show eigenvalues exist, and to prove Schur's lemma. +lemma nonempty_of_is_alg_closed_of_finite_dimensional [is_alg_closed 𝕜] + [nontrivial A] [I : finite_dimensional 𝕜 A] (a : A) : + (σ a).nonempty := +begin + obtain ⟨p, ⟨h_mon, h_eval_p⟩⟩ := is_integral_of_noetherian (is_noetherian.iff_fg.2 I) a, + have nu : ¬ is_unit (aeval a p), { rw [←aeval_def] at h_eval_p, rw h_eval_p, simp, }, + rw [eq_prod_roots_of_monic_of_splits_id h_mon (is_alg_closed.splits p)] at nu, + obtain ⟨k, hk, _⟩ := exists_mem_of_not_is_unit_aeval_prod (monic.ne_zero h_mon) nu, + exact ⟨k, hk⟩ +end + +end scalar_field + +end spectrum diff --git a/src/linear_algebra/eigenspace.lean b/src/linear_algebra/eigenspace.lean index 1ea0617830ae5..2d18823e0749b 100644 --- a/src/linear_algebra/eigenspace.lean +++ b/src/linear_algebra/eigenspace.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Alexander Bentkamp -/ -import algebra.algebra.spectrum +import field_theory.is_alg_closed.spectrum import order.hom.basic import linear_algebra.general_linear_group From c20927220ef87bb4962ba08bf6da2ce3cf50a6dd Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Wed, 7 Jun 2023 05:24:06 +0000 Subject: [PATCH 1150/1291] chore(*): add mathlib4 synchronization comments (#19162) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.category.Module.colimits` * `algebra.category.Module.monoidal.closed` * `algebra.homology.local_cohomology` * `analysis.box_integral.basic` * `analysis.box_integral.divergence_theorem` * `analysis.box_integral.integrability` * `analysis.calculus.parametric_integral` * `analysis.convex.integral` * `analysis.convex.specific_functions.deriv` * `analysis.special_functions.exponential` * `analysis.special_functions.pow.deriv` * `analysis.special_functions.trigonometric.arctan` * `analysis.special_functions.trigonometric.complex` * `analysis.special_functions.trigonometric.complex_deriv` * `analysis.special_functions.trigonometric.series` * `category_theory.closed.functor_category` * `category_theory.with_terminal` * `combinatorics.additive.behrend` * `combinatorics.derangements.exponential` * `data.set.ncard` * `field_theory.splitting_field.is_splitting_field` * `group_theory.schreier` * `group_theory.specific_groups.quaternion` * `linear_algebra.quadratic_form.complex` * `linear_algebra.quadratic_form.isometry` * `linear_algebra.quadratic_form.prod` * `linear_algebra.quadratic_form.real` * `measure_theory.constructions.prod.integral` * `measure_theory.function.ae_eq_of_integral` * `measure_theory.group.fundamental_domain` * `measure_theory.integral.average` * `measure_theory.integral.set_integral` * `measure_theory.measure.haar.inner_product_space` * `measure_theory.measure.with_density_vector_measure` * `representation_theory.Action` * `ring_theory.dedekind_domain.factorization` * `ring_theory.localization.away.adjoin_root` * `ring_theory.valuation.valuation_ring` --- src/algebra/category/Module/colimits.lean | 3 +++ src/algebra/category/Module/monoidal/closed.lean | 3 +++ src/algebra/homology/local_cohomology.lean | 3 +++ src/analysis/box_integral/basic.lean | 3 +++ src/analysis/box_integral/divergence_theorem.lean | 3 +++ src/analysis/box_integral/integrability.lean | 3 +++ src/analysis/calculus/parametric_integral.lean | 3 +++ src/analysis/convex/integral.lean | 3 +++ src/analysis/convex/specific_functions/deriv.lean | 3 +++ src/analysis/special_functions/exponential.lean | 3 +++ src/analysis/special_functions/pow/deriv.lean | 3 +++ src/analysis/special_functions/trigonometric/arctan.lean | 3 +++ src/analysis/special_functions/trigonometric/complex.lean | 3 +++ .../special_functions/trigonometric/complex_deriv.lean | 3 +++ src/analysis/special_functions/trigonometric/series.lean | 3 +++ src/category_theory/closed/functor_category.lean | 3 +++ src/category_theory/with_terminal.lean | 3 +++ src/combinatorics/additive/behrend.lean | 3 +++ src/combinatorics/derangements/exponential.lean | 3 +++ src/data/set/ncard.lean | 3 +++ src/field_theory/splitting_field/is_splitting_field.lean | 3 +++ src/group_theory/schreier.lean | 3 +++ src/group_theory/specific_groups/quaternion.lean | 3 +++ src/linear_algebra/quadratic_form/complex.lean | 3 +++ src/linear_algebra/quadratic_form/isometry.lean | 3 +++ src/linear_algebra/quadratic_form/prod.lean | 3 +++ src/linear_algebra/quadratic_form/real.lean | 3 +++ src/measure_theory/constructions/prod/integral.lean | 3 +++ src/measure_theory/function/ae_eq_of_integral.lean | 3 +++ src/measure_theory/group/fundamental_domain.lean | 3 +++ src/measure_theory/integral/average.lean | 3 +++ src/measure_theory/integral/set_integral.lean | 3 +++ src/measure_theory/measure/haar/inner_product_space.lean | 3 +++ src/measure_theory/measure/with_density_vector_measure.lean | 3 +++ src/representation_theory/Action.lean | 3 +++ src/ring_theory/dedekind_domain/factorization.lean | 3 +++ src/ring_theory/localization/away/adjoin_root.lean | 3 +++ src/ring_theory/valuation/valuation_ring.lean | 3 +++ 38 files changed, 114 insertions(+) diff --git a/src/algebra/category/Module/colimits.lean b/src/algebra/category/Module/colimits.lean index eaca92284e5c7..1bc29635f615e 100644 --- a/src/algebra/category/Module/colimits.lean +++ b/src/algebra/category/Module/colimits.lean @@ -9,6 +9,9 @@ import category_theory.concrete_category.elementwise /-! # The category of R-modules has all colimits. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file uses a "pre-automated" approach, just as for `Mon/colimits.lean`. Note that finite colimits can already be obtained from the instance `abelian (Module R)`. diff --git a/src/algebra/category/Module/monoidal/closed.lean b/src/algebra/category/Module/monoidal/closed.lean index 498cff23a9b50..1f2577cf89c0f 100644 --- a/src/algebra/category/Module/monoidal/closed.lean +++ b/src/algebra/category/Module/monoidal/closed.lean @@ -8,6 +8,9 @@ import algebra.category.Module.monoidal.symmetric /-! # The monoidal closed structure on `Module R`. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ universes v w x u diff --git a/src/algebra/homology/local_cohomology.lean b/src/algebra/homology/local_cohomology.lean index ec568744815d9..ea2ccd6ee952a 100644 --- a/src/algebra/homology/local_cohomology.lean +++ b/src/algebra/homology/local_cohomology.lean @@ -12,6 +12,9 @@ import ring_theory.finiteness /-! # Local cohomology. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the `i`-th local cohomology module of an `R`-module `M` with support in an ideal `I` of `R`, where `R` is a commutative ring, as the direct limit of Ext modules: diff --git a/src/analysis/box_integral/basic.lean b/src/analysis/box_integral/basic.lean index eb617857d3890..a02b18afc77bd 100644 --- a/src/analysis/box_integral/basic.lean +++ b/src/analysis/box_integral/basic.lean @@ -10,6 +10,9 @@ import topology.uniform_space.compact /-! # Integrals of Riemann, Henstock-Kurzweil, and McShane +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define the integral of a function over a box in `ℝⁿ. The same definition works for Riemann, Henstock-Kurzweil, and McShane integrals. diff --git a/src/analysis/box_integral/divergence_theorem.lean b/src/analysis/box_integral/divergence_theorem.lean index 14cdbba22fc83..4744a0364b2b8 100644 --- a/src/analysis/box_integral/divergence_theorem.lean +++ b/src/analysis/box_integral/divergence_theorem.lean @@ -13,6 +13,9 @@ import analysis.calculus.fderiv.restrict_scalars /-! # Divergence integral for Henstock-Kurzweil integral +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove the Divergence Theorem for a Henstock-Kurzweil style integral. The theorem says the following. Let `f : ℝⁿ → Eⁿ` be a function differentiable on a closed rectangular box `I` with derivative `f' x : ℝⁿ →L[ℝ] Eⁿ` at `x ∈ I`. Then the divergence `λ x, ∑ k, f' x eₖ k`, diff --git a/src/analysis/box_integral/integrability.lean b/src/analysis/box_integral/integrability.lean index bc9d7862b4ea9..6a1630f6cd8a3 100644 --- a/src/analysis/box_integral/integrability.lean +++ b/src/analysis/box_integral/integrability.lean @@ -10,6 +10,9 @@ import measure_theory.measure.regular /-! # McShane integrability vs Bochner integrability +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove that any Bochner integrable function is McShane integrable (hence, it is Henstock and `⊥` integrable) with the same integral. The proof is based on [Russel A. Gordon, *The integrals of Lebesgue, Denjoy, Perron, and Henstock*][Gordon55]. diff --git a/src/analysis/calculus/parametric_integral.lean b/src/analysis/calculus/parametric_integral.lean index 27b7e59312900..307f254a58143 100644 --- a/src/analysis/calculus/parametric_integral.lean +++ b/src/analysis/calculus/parametric_integral.lean @@ -9,6 +9,9 @@ import measure_theory.integral.set_integral /-! # Derivatives of integrals depending on parameters +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A parametric integral is a function with shape `f = λ x : H, ∫ a : α, F x a ∂μ` for some `F : H → α → E`, where `H` and `E` are normed spaces and `α` is a measured space with measure `μ`. diff --git a/src/analysis/convex/integral.lean b/src/analysis/convex/integral.lean index 2ae4b1cb63efc..11d6ab7a75741 100644 --- a/src/analysis/convex/integral.lean +++ b/src/analysis/convex/integral.lean @@ -11,6 +11,9 @@ import measure_theory.integral.average /-! # Jensen's inequality for integrals +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove several forms of Jensen's inequality for integrals. - for convex sets: `convex.average_mem`, `convex.set_average_mem`, `convex.integral_mem`; diff --git a/src/analysis/convex/specific_functions/deriv.lean b/src/analysis/convex/specific_functions/deriv.lean index f8ed93d429deb..60aa49399d058 100644 --- a/src/analysis/convex/specific_functions/deriv.lean +++ b/src/analysis/convex/specific_functions/deriv.lean @@ -10,6 +10,9 @@ import analysis.special_functions.sqrt /-! # Collection of convex functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove that certain specific functions are strictly convex, including the following: * `even.strict_convex_on_pow` : For an even `n : ℕ` with `2 ≤ n`, `λ x, x ^ n` is strictly convex. diff --git a/src/analysis/special_functions/exponential.lean b/src/analysis/special_functions/exponential.lean index 3208685d55087..f928d78a10f14 100644 --- a/src/analysis/special_functions/exponential.lean +++ b/src/analysis/special_functions/exponential.lean @@ -10,6 +10,9 @@ import topology.metric_space.cau_seq_filter /-! # Calculus results on exponential in a Banach algebra +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we prove basic properties about the derivative of the exponential map `exp 𝕂` in a Banach algebra `𝔸` over a field `𝕂`. We keep them separate from the main file `analysis/normed_space/exponential` in order to minimize dependencies. diff --git a/src/analysis/special_functions/pow/deriv.lean b/src/analysis/special_functions/pow/deriv.lean index ee55bb82b775d..a65165d8bd4ef 100644 --- a/src/analysis/special_functions/pow/deriv.lean +++ b/src/analysis/special_functions/pow/deriv.lean @@ -14,6 +14,9 @@ import analysis.special_functions.trigonometric.deriv /-! # Derivatives of power function on `ℂ`, `ℝ`, `ℝ≥0`, and `ℝ≥0∞` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We also prove differentiability and provide derivatives for the power functions `x ^ y`. -/ diff --git a/src/analysis/special_functions/trigonometric/arctan.lean b/src/analysis/special_functions/trigonometric/arctan.lean index 10ce7f424c9fc..1b00d5a47d1b4 100644 --- a/src/analysis/special_functions/trigonometric/arctan.lean +++ b/src/analysis/special_functions/trigonometric/arctan.lean @@ -8,6 +8,9 @@ import analysis.special_functions.trigonometric.complex /-! # The `arctan` function. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Inequalities, derivatives, and `real.tan` as a `local_homeomorph` between `(-(π / 2), π / 2)` and the whole line. -/ diff --git a/src/analysis/special_functions/trigonometric/complex.lean b/src/analysis/special_functions/trigonometric/complex.lean index 35a037e540bd6..ddcc993a29a5f 100644 --- a/src/analysis/special_functions/trigonometric/complex.lean +++ b/src/analysis/special_functions/trigonometric/complex.lean @@ -9,6 +9,9 @@ import analysis.convex.specific_functions.deriv /-! # Complex trigonometric functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Basic facts and derivatives for the complex trigonometric functions. Several facts about the real trigonometric functions have the proofs deferred here, rather than diff --git a/src/analysis/special_functions/trigonometric/complex_deriv.lean b/src/analysis/special_functions/trigonometric/complex_deriv.lean index a1c2f4550f6e8..2ab19256c1269 100644 --- a/src/analysis/special_functions/trigonometric/complex_deriv.lean +++ b/src/analysis/special_functions/trigonometric/complex_deriv.lean @@ -8,6 +8,9 @@ import analysis.special_functions.trigonometric.complex /-! # Complex trigonometric functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Basic facts and derivatives for the complex trigonometric functions. -/ diff --git a/src/analysis/special_functions/trigonometric/series.lean b/src/analysis/special_functions/trigonometric/series.lean index 0665bc3e9122f..a3ece09d3d75a 100644 --- a/src/analysis/special_functions/trigonometric/series.lean +++ b/src/analysis/special_functions/trigonometric/series.lean @@ -7,6 +7,9 @@ import analysis.special_functions.exponential /-! # Trigonometric functions as sums of infinite series +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we express trigonometric functions in terms of their series expansion. ## Main results diff --git a/src/category_theory/closed/functor_category.lean b/src/category_theory/closed/functor_category.lean index c6633e2f078aa..029c5c26d1c1d 100644 --- a/src/category_theory/closed/functor_category.lean +++ b/src/category_theory/closed/functor_category.lean @@ -9,6 +9,9 @@ import category_theory.monoidal.functor_category /-! # Functors from a groupoid into a monoidal closed category form a monoidal closed category. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + (Using the pointwise monoidal structure on the functor category.) -/ diff --git a/src/category_theory/with_terminal.lean b/src/category_theory/with_terminal.lean index 28b0f21f32763..cab0ee066e006 100644 --- a/src/category_theory/with_terminal.lean +++ b/src/category_theory/with_terminal.lean @@ -9,6 +9,9 @@ import category_theory.limits.shapes.terminal # `with_initial` and `with_terminal` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given a category `C`, this file constructs two objects: 1. `with_terminal C`, the category built from `C` by formally adjoining a terminal object. 2. `with_initial C`, the category built from `C` by formally adjoining an initial object. diff --git a/src/combinatorics/additive/behrend.lean b/src/combinatorics/additive/behrend.lean index 857a123b6c14e..1257399320c11 100644 --- a/src/combinatorics/additive/behrend.lean +++ b/src/combinatorics/additive/behrend.lean @@ -11,6 +11,9 @@ import data.complex.exponential_bounds /-! # Behrend's bound on Roth numbers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves Behrend's lower bound on Roth numbers. This says that we can find a subset of `{1, ..., n}` of size `n / exp (O (sqrt (log n)))` which does not contain arithmetic progressions of length `3`. diff --git a/src/combinatorics/derangements/exponential.lean b/src/combinatorics/derangements/exponential.lean index 6ab726aa8a1ce..4bb378f6b3127 100644 --- a/src/combinatorics/derangements/exponential.lean +++ b/src/combinatorics/derangements/exponential.lean @@ -10,6 +10,9 @@ import order.filter.basic /-! # Derangement exponential series +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves that the probability of a permutation on n elements being a derangement is 1/e. The specific lemma is `num_derangements_tendsto_inv_e`. -/ diff --git a/src/data/set/ncard.lean b/src/data/set/ncard.lean index 965ef766b1616..34ccdf0fa368a 100644 --- a/src/data/set/ncard.lean +++ b/src/data/set/ncard.lean @@ -9,6 +9,9 @@ import algebra.big_operators.finprod /-! # Noncomputable Set Cardinality +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define the cardinality `set.ncard s` of a set `s` as a natural number. This function is noncomputable (being defined in terms of `nat.card`) and takes the value `0` if `s` is infinite. diff --git a/src/field_theory/splitting_field/is_splitting_field.lean b/src/field_theory/splitting_field/is_splitting_field.lean index fd3dbd83e6d0d..8d93b01342067 100644 --- a/src/field_theory/splitting_field/is_splitting_field.lean +++ b/src/field_theory/splitting_field/is_splitting_field.lean @@ -10,6 +10,9 @@ import ring_theory.adjoin.field /-! # Splitting fields +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file introduces the notion of a splitting field of a polynomial and provides an embedding from a splitting field to any field that splits the polynomial. A polynomial `f : K[X]` splits over a field extension `L` of `K` if it is zero or all of its irreducible factors over `L` have diff --git a/src/group_theory/schreier.lean b/src/group_theory/schreier.lean index 8cc1440ee1d67..b13b3323c061d 100644 --- a/src/group_theory/schreier.lean +++ b/src/group_theory/schreier.lean @@ -11,6 +11,9 @@ import group_theory.transfer /-! # Schreier's Lemma +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove Schreier's lemma. ## Main results diff --git a/src/group_theory/specific_groups/quaternion.lean b/src/group_theory/specific_groups/quaternion.lean index fac7aa55f9907..0af5ef908c007 100644 --- a/src/group_theory/specific_groups/quaternion.lean +++ b/src/group_theory/specific_groups/quaternion.lean @@ -12,6 +12,9 @@ import group_theory.specific_groups.cyclic /-! # Quaternion Groups +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define the (generalised) quaternion groups `quaternion_group n` of order `4n`, also known as dicyclic groups, with elements `a i` and `xa i` for `i : zmod n`. The (generalised) quaternion groups can be defined by the presentation diff --git a/src/linear_algebra/quadratic_form/complex.lean b/src/linear_algebra/quadratic_form/complex.lean index 8d7c5c33434fb..9a42c8445a840 100644 --- a/src/linear_algebra/quadratic_form/complex.lean +++ b/src/linear_algebra/quadratic_form/complex.lean @@ -9,6 +9,9 @@ import analysis.special_functions.pow.complex /-! # Quadratic forms over the complex numbers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + `equivalent_sum_squares`: A nondegenerate quadratic form over the complex numbers is equivalent to a sum of squares. diff --git a/src/linear_algebra/quadratic_form/isometry.lean b/src/linear_algebra/quadratic_form/isometry.lean index 36af1c374f979..7c97543e99937 100644 --- a/src/linear_algebra/quadratic_form/isometry.lean +++ b/src/linear_algebra/quadratic_form/isometry.lean @@ -9,6 +9,9 @@ import linear_algebra.quadratic_form.basic /-! # Isometries with respect to quadratic forms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions * `quadratic_form.isometry`: `linear_equiv`s which map between two different quadratic forms diff --git a/src/linear_algebra/quadratic_form/prod.lean b/src/linear_algebra/quadratic_form/prod.lean index 2f8ee90aeff7f..5b7d0039b76c5 100644 --- a/src/linear_algebra/quadratic_form/prod.lean +++ b/src/linear_algebra/quadratic_form/prod.lean @@ -7,6 +7,9 @@ import linear_algebra.quadratic_form.isometry /-! # Quadratic form on product and pi types +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions * `quadratic_form.prod Q₁ Q₂`: the quadratic form constructed elementwise on a product diff --git a/src/linear_algebra/quadratic_form/real.lean b/src/linear_algebra/quadratic_form/real.lean index 289a2b2b2b552..f3ae98248de69 100644 --- a/src/linear_algebra/quadratic_form/real.lean +++ b/src/linear_algebra/quadratic_form/real.lean @@ -10,6 +10,9 @@ import data.real.sign /-! # Real quadratic forms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Sylvester's law of inertia `equivalent_one_neg_one_weighted_sum_squared`: A real quadratic form is equivalent to a weighted sum of squares with the weights being ±1 or 0. diff --git a/src/measure_theory/constructions/prod/integral.lean b/src/measure_theory/constructions/prod/integral.lean index dab3b0377a83b..e92c1173627f9 100644 --- a/src/measure_theory/constructions/prod/integral.lean +++ b/src/measure_theory/constructions/prod/integral.lean @@ -9,6 +9,9 @@ import measure_theory.integral.set_integral /-! # Integration with respect to the product measure +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove Fubini's theorem. ## Main results diff --git a/src/measure_theory/function/ae_eq_of_integral.lean b/src/measure_theory/function/ae_eq_of_integral.lean index 6358a66f7fa7d..27c711e645a7e 100644 --- a/src/measure_theory/function/ae_eq_of_integral.lean +++ b/src/measure_theory/function/ae_eq_of_integral.lean @@ -11,6 +11,9 @@ import measure_theory.integral.set_integral /-! # From equality of integrals to equality of functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file provides various statements of the general form "if two functions have the same integral on all sets, then they are equal almost everywhere". The different lemmas use various hypotheses on the class of functions, on the target space or on the diff --git a/src/measure_theory/group/fundamental_domain.lean b/src/measure_theory/group/fundamental_domain.lean index 689cacb93d810..ec47f9ce9daf4 100644 --- a/src/measure_theory/group/fundamental_domain.lean +++ b/src/measure_theory/group/fundamental_domain.lean @@ -9,6 +9,9 @@ import measure_theory.integral.set_integral /-! # Fundamental domain of a group action +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A set `s` is said to be a *fundamental domain* of an action of a group `G` on a measurable space `α` with respect to a measure `μ` if diff --git a/src/measure_theory/integral/average.lean b/src/measure_theory/integral/average.lean index e11fa52b50090..efd336ea1c821 100644 --- a/src/measure_theory/integral/average.lean +++ b/src/measure_theory/integral/average.lean @@ -8,6 +8,9 @@ import measure_theory.integral.set_integral /-! # Integral average of a function +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define `measure_theory.average μ f` (notation: `⨍ x, f x ∂μ`) to be the average value of `f` with respect to measure `μ`. It is defined as `∫ x, f x ∂((μ univ)⁻¹ • μ)`, so it is equal to zero if `f` is not integrable or if `μ` is an infinite measure. If `μ` is a probability diff --git a/src/measure_theory/integral/set_integral.lean b/src/measure_theory/integral/set_integral.lean index e1754e0df4e65..a4ef51091adbb 100644 --- a/src/measure_theory/integral/set_integral.lean +++ b/src/measure_theory/integral/set_integral.lean @@ -13,6 +13,9 @@ import topology.continuous_function.compact /-! # Set integral +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove some properties of `∫ x in s, f x ∂μ`. Recall that this notation is defined as `∫ x, f x ∂(μ.restrict s)`. In `integral_indicator` we prove that for a measurable function `f` and a measurable set `s` this definition coincides with another natural definition: diff --git a/src/measure_theory/measure/haar/inner_product_space.lean b/src/measure_theory/measure/haar/inner_product_space.lean index 999edc737235e..8d7fed18e4f47 100644 --- a/src/measure_theory/measure/haar/inner_product_space.lean +++ b/src/measure_theory/measure/haar/inner_product_space.lean @@ -9,6 +9,9 @@ import measure_theory.measure.lebesgue.eq_haar /-! # Volume forms and measures on inner product spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A volume form induces a Lebesgue measure on general finite-dimensional real vector spaces. In this file, we discuss the specific situation of inner product spaces, where an orientation gives rise to a canonical volume form. We show that the measure coming from this volume form gives diff --git a/src/measure_theory/measure/with_density_vector_measure.lean b/src/measure_theory/measure/with_density_vector_measure.lean index fb74f73525300..20d61de0cff09 100644 --- a/src/measure_theory/measure/with_density_vector_measure.lean +++ b/src/measure_theory/measure/with_density_vector_measure.lean @@ -10,6 +10,9 @@ import measure_theory.function.ae_eq_of_integral # Vector measure defined by an integral +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given a measure `μ` and an integrable function `f : α → E`, we can define a vector measure `v` such that for all measurable set `s`, `v i = ∫ x in s, f x ∂μ`. This definition is useful for the Radon-Nikodym theorem for signed measures. diff --git a/src/representation_theory/Action.lean b/src/representation_theory/Action.lean index 2a6fc140cb831..df29ca63dca5a 100644 --- a/src/representation_theory/Action.lean +++ b/src/representation_theory/Action.lean @@ -23,6 +23,9 @@ import category_theory.linear.functor_category /-! # `Action V G`, the category of actions of a monoid `G` inside some category `V`. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The prototypical example is `V = Module R`, where `Action (Module R) G` is the category of `R`-linear representations of `G`. diff --git a/src/ring_theory/dedekind_domain/factorization.lean b/src/ring_theory/dedekind_domain/factorization.lean index 3527aeb7fcd0b..ad23edb9a51f3 100644 --- a/src/ring_theory/dedekind_domain/factorization.lean +++ b/src/ring_theory/dedekind_domain/factorization.lean @@ -6,6 +6,9 @@ Authors: María Inés de Frutos-Fernández import ring_theory.dedekind_domain.ideal /-! # Factorization of ideals of Dedekind domains + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. Every nonzero ideal `I` of a Dedekind domain `R` can be factored as a product `∏_v v^{n_v}` over the maximal ideals of `R`, where the exponents `n_v` are natural numbers. TODO: Extend the results in this file to fractional ideals of `R`. diff --git a/src/ring_theory/localization/away/adjoin_root.lean b/src/ring_theory/localization/away/adjoin_root.lean index c169a02f9172c..2d5ed60a9769c 100644 --- a/src/ring_theory/localization/away/adjoin_root.lean +++ b/src/ring_theory/localization/away/adjoin_root.lean @@ -7,6 +7,9 @@ import ring_theory.adjoin_root import ring_theory.localization.away.basic /-! +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The `R`-`alg_equiv` between the localization of `R` away from `r` and `R` with an inverse of `r` adjoined. -/ diff --git a/src/ring_theory/valuation/valuation_ring.lean b/src/ring_theory/valuation/valuation_ring.lean index e52900b6e6bfa..9cc608de0f3bb 100644 --- a/src/ring_theory/valuation/valuation_ring.lean +++ b/src/ring_theory/valuation/valuation_ring.lean @@ -14,6 +14,9 @@ import tactic.field_simp /-! # Valuation Rings +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A valuation ring is a domain such that for every pair of elements `a b`, either `a` divides `b` or vice-versa. From 6b0169218d01f2837d79ea2784882009a0da1aa1 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Wed, 7 Jun 2023 06:47:08 +0000 Subject: [PATCH 1151/1291] chore(linear_algebra/eigenspace): split file (#19163) Co-authored-by: Scott Morrison --- src/algebra/lie/nilpotent.lean | 2 +- src/algebra/lie/weights.lean | 2 +- .../inner_product_space/rayleigh.lean | 2 +- .../inner_product_space/spectrum.lean | 1 + .../basic.lean} | 182 ++---------------- .../eigenspace/is_alg_closed.lean | 115 +++++++++++ src/linear_algebra/eigenspace/minpoly.lean | 113 +++++++++++ 7 files changed, 243 insertions(+), 174 deletions(-) rename src/linear_algebra/{eigenspace.lean => eigenspace/basic.lean} (73%) create mode 100644 src/linear_algebra/eigenspace/is_alg_closed.lean create mode 100644 src/linear_algebra/eigenspace/minpoly.lean diff --git a/src/algebra/lie/nilpotent.lean b/src/algebra/lie/nilpotent.lean index 3f7c43ddd61d1..1f614879622e9 100644 --- a/src/algebra/lie/nilpotent.lean +++ b/src/algebra/lie/nilpotent.lean @@ -6,7 +6,7 @@ Authors: Oliver Nash import algebra.lie.solvable import algebra.lie.quotient import algebra.lie.normalizer -import linear_algebra.eigenspace +import linear_algebra.eigenspace.basic import ring_theory.nilpotent /-! diff --git a/src/algebra/lie/weights.lean b/src/algebra/lie/weights.lean index a2a68d870af65..cfaced0c55890 100644 --- a/src/algebra/lie/weights.lean +++ b/src/algebra/lie/weights.lean @@ -8,7 +8,7 @@ import algebra.lie.tensor_product import algebra.lie.character import algebra.lie.engel import algebra.lie.cartan_subalgebra -import linear_algebra.eigenspace +import linear_algebra.eigenspace.basic import ring_theory.tensor_product /-! diff --git a/src/analysis/inner_product_space/rayleigh.lean b/src/analysis/inner_product_space/rayleigh.lean index 4fc947239e77d..a01efa5925ee3 100644 --- a/src/analysis/inner_product_space/rayleigh.lean +++ b/src/analysis/inner_product_space/rayleigh.lean @@ -7,7 +7,7 @@ import analysis.inner_product_space.calculus import analysis.inner_product_space.dual import analysis.inner_product_space.adjoint import analysis.calculus.lagrange_multipliers -import linear_algebra.eigenspace +import linear_algebra.eigenspace.basic /-! # The Rayleigh quotient diff --git a/src/analysis/inner_product_space/spectrum.lean b/src/analysis/inner_product_space/spectrum.lean index 5d13cff964c7e..d3a9f2866d8e8 100644 --- a/src/analysis/inner_product_space/spectrum.lean +++ b/src/analysis/inner_product_space/spectrum.lean @@ -6,6 +6,7 @@ Authors: Heather Macbeth import analysis.inner_product_space.rayleigh import analysis.inner_product_space.pi_L2 import algebra.direct_sum.decomposition +import linear_algebra.eigenspace.minpoly /-! # Spectral theory of self-adjoint operators diff --git a/src/linear_algebra/eigenspace.lean b/src/linear_algebra/eigenspace/basic.lean similarity index 73% rename from src/linear_algebra/eigenspace.lean rename to src/linear_algebra/eigenspace/basic.lean index 2d18823e0749b..fcfe2caf3a4bf 100644 --- a/src/linear_algebra/eigenspace.lean +++ b/src/linear_algebra/eigenspace/basic.lean @@ -4,9 +4,10 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Alexander Bentkamp -/ -import field_theory.is_alg_closed.spectrum -import order.hom.basic +import algebra.algebra.spectrum import linear_algebra.general_linear_group +import linear_algebra.finite_dimensional + /-! # Eigenvectors and eigenvalues @@ -28,6 +29,13 @@ of the map `(f - μ • id) ^ k`. The nonzero elements of a generalized eigenspa eigenvectors `x`. If there are generalized eigenvectors for a natural number `k` and a scalar `μ`, the scalar `μ` is called a generalized eigenvalue. +The fact that the eigenvalues are the roots of the minimal polynomial is proved in +`linear_algebra.eigenspace.minpoly`. + +The existence of eigenvalues over an algebraically closed field +(and the fact that the generalized eigenspaces then span) is deferred to +`linear_algebra.eigenspace.is_alg_closed`. + ## References * [Sheldon Axler, *Linear Algebra Done Right*][axler2015] @@ -43,8 +51,7 @@ universes u v w namespace module namespace End -open module principal_ideal_ring polynomial finite_dimensional -open_locale polynomial +open finite_dimensional variables {K R : Type v} {V M : Type w} [comm_ring R] [add_comm_group M] [module R M] [field K] [add_comm_group V] [module K V] @@ -115,106 +122,6 @@ calc ... = (b • (f - b⁻¹ • algebra_map K (End K V) a)).ker : by rw linear_map.ker_smul _ b hb ... = (b • f - algebra_map K (End K V) a).ker : by rw [smul_sub, smul_inv_smul₀ hb] -lemma eigenspace_aeval_polynomial_degree_1 - (f : End K V) (q : K[X]) (hq : degree q = 1) : - eigenspace f (- q.coeff 0 / q.leading_coeff) = (aeval f q).ker := -calc - eigenspace f (- q.coeff 0 / q.leading_coeff) - = (q.leading_coeff • f - algebra_map K (End K V) (- q.coeff 0)).ker - : by { rw eigenspace_div, intro h, rw leading_coeff_eq_zero_iff_deg_eq_bot.1 h at hq, cases hq } - ... = (aeval f (C q.leading_coeff * X + C (q.coeff 0))).ker - : by { rw [C_mul', aeval_def], simp [algebra_map, algebra.to_ring_hom], } - ... = (aeval f q).ker - : by rwa ← eq_X_add_C_of_degree_eq_one - -lemma ker_aeval_ring_hom'_unit_polynomial - (f : End K V) (c : (K[X])ˣ) : - (aeval f (c : K[X])).ker = ⊥ := -begin - rw polynomial.eq_C_of_degree_eq_zero (degree_coe_units c), - simp only [aeval_def, eval₂_C], - apply ker_algebra_map_End, - apply coeff_coe_units_zero_ne_zero c -end - -theorem aeval_apply_of_has_eigenvector {f : End K V} - {p : K[X]} {μ : K} {x : V} (h : f.has_eigenvector μ x) : - aeval f p x = (p.eval μ) • x := -begin - apply p.induction_on, - { intro a, simp [module.algebra_map_End_apply] }, - { intros p q hp hq, simp [hp, hq, add_smul] }, - { intros n a hna, - rw [mul_comm, pow_succ, mul_assoc, alg_hom.map_mul, linear_map.mul_apply, mul_comm, hna], - simp only [mem_eigenspace_iff.1 h.1, smul_smul, aeval_X, eval_mul, eval_C, eval_pow, eval_X, - linear_map.map_smulₛₗ, ring_hom.id_apply, mul_comm] } -end - -section minpoly - -theorem is_root_of_has_eigenvalue {f : End K V} {μ : K} (h : f.has_eigenvalue μ) : - (minpoly K f).is_root μ := -begin - rcases (submodule.ne_bot_iff _).1 h with ⟨w, ⟨H, ne0⟩⟩, - refine or.resolve_right (smul_eq_zero.1 _) ne0, - simp [← aeval_apply_of_has_eigenvector ⟨H, ne0⟩, minpoly.aeval K f], -end - -variables [finite_dimensional K V] (f : End K V) - -variables {f} {μ : K} - -theorem has_eigenvalue_of_is_root (h : (minpoly K f).is_root μ) : - f.has_eigenvalue μ := -begin - cases dvd_iff_is_root.2 h with p hp, - rw [has_eigenvalue, eigenspace], - intro con, - cases (linear_map.is_unit_iff_ker_eq_bot _).2 con with u hu, - have p_ne_0 : p ≠ 0, - { intro con, - apply minpoly.ne_zero f.is_integral, - rw [hp, con, mul_zero] }, - have h_deg := minpoly.degree_le_of_ne_zero K f p_ne_0 _, - { rw [hp, degree_mul, degree_X_sub_C, polynomial.degree_eq_nat_degree p_ne_0] at h_deg, - norm_cast at h_deg, - linarith, }, - { have h_aeval := minpoly.aeval K f, - revert h_aeval, - simp [hp, ← hu] }, -end - -theorem has_eigenvalue_iff_is_root : - f.has_eigenvalue μ ↔ (minpoly K f).is_root μ := -⟨is_root_of_has_eigenvalue, has_eigenvalue_of_is_root⟩ - -/-- An endomorphism of a finite-dimensional vector space has finitely many eigenvalues. -/ -noncomputable instance (f : End K V) : fintype f.eigenvalues := -set.finite.fintype -begin - have h : minpoly K f ≠ 0 := minpoly.ne_zero f.is_integral, - convert (minpoly K f).root_set_finite K, - ext μ, - have : (μ ∈ {μ : K | f.eigenspace μ = ⊥ → false}) ↔ ¬f.eigenspace μ = ⊥ := by tauto, - convert rfl.mpr this, - simp [polynomial.root_set_def, polynomial.mem_roots h, ← has_eigenvalue_iff_is_root, - has_eigenvalue] -end - -end minpoly - -/-- Every linear operator on a vector space over an algebraically closed field has - an eigenvalue. -/ --- This is Lemma 5.21 of [axler2015], although we are no longer following that proof. -lemma exists_eigenvalue [is_alg_closed K] [finite_dimensional K V] [nontrivial V] (f : End K V) : - ∃ (c : K), f.has_eigenvalue c := -by { simp_rw has_eigenvalue_iff_mem_spectrum, - exact spectrum.nonempty_of_is_alg_closed_of_finite_dimensional K f } - -noncomputable instance [is_alg_closed K] [finite_dimensional K V] [nontrivial V] (f : End K V) : - inhabited f.eigenvalues := -⟨⟨f.exists_eigenvalue.some, f.exists_eigenvalue.some_spec⟩⟩ - /-- The eigenspaces of a linear operator form an independent family of subspaces of `V`. That is, any eigenspace has trivial intersection with the span of all the other eigenspaces. -/ lemma eigenspaces_independent (f : End K V) : complete_lattice.independent f.eigenspace := @@ -530,72 +437,5 @@ calc submodule.map f (f.generalized_eigenrange μ n) ... = submodule.map ((f - algebra_map _ _ μ) ^ n) f.range : linear_map.range_comp _ _ ... ≤ f.generalized_eigenrange μ n : linear_map.map_le_range -/-- The generalized eigenvectors span the entire vector space (Lemma 8.21 of [axler2015]). -/ -lemma supr_generalized_eigenspace_eq_top [is_alg_closed K] [finite_dimensional K V] (f : End K V) : - (⨆ (μ : K) (k : ℕ), f.generalized_eigenspace μ k) = ⊤ := -begin - -- We prove the claim by strong induction on the dimension of the vector space. - unfreezingI { induction h_dim : finrank K V using nat.strong_induction_on - with n ih generalizing V }, - cases n, - -- If the vector space is 0-dimensional, the result is trivial. - { rw ←top_le_iff, - simp only [finrank_eq_zero.1 (eq.trans (finrank_top _ _) h_dim), bot_le] }, - -- Otherwise the vector space is nontrivial. - { haveI : nontrivial V := finrank_pos_iff.1 (by { rw h_dim, apply nat.zero_lt_succ }), - -- Hence, `f` has an eigenvalue `μ₀`. - obtain ⟨μ₀, hμ₀⟩ : ∃ μ₀, f.has_eigenvalue μ₀ := exists_eigenvalue f, - -- We define `ES` to be the generalized eigenspace - let ES := f.generalized_eigenspace μ₀ (finrank K V), - -- and `ER` to be the generalized eigenrange. - let ER := f.generalized_eigenrange μ₀ (finrank K V), - -- `f` maps `ER` into itself. - have h_f_ER : ∀ (x : V), x ∈ ER → f x ∈ ER, - from λ x hx, map_generalized_eigenrange_le (submodule.mem_map_of_mem hx), - -- Therefore, we can define the restriction `f'` of `f` to `ER`. - let f' : End K ER := f.restrict h_f_ER, - -- The dimension of `ES` is positive - have h_dim_ES_pos : 0 < finrank K ES, - { dsimp only [ES], - rw h_dim, - apply pos_finrank_generalized_eigenspace_of_has_eigenvalue hμ₀ (nat.zero_lt_succ n) }, - -- and the dimensions of `ES` and `ER` add up to `finrank K V`. - have h_dim_add : finrank K ER + finrank K ES = finrank K V, - { apply linear_map.finrank_range_add_finrank_ker }, - -- Therefore the dimension `ER` mus be smaller than `finrank K V`. - have h_dim_ER : finrank K ER < n.succ, by linarith, - -- This allows us to apply the induction hypothesis on `ER`: - have ih_ER : (⨆ (μ : K) (k : ℕ), f'.generalized_eigenspace μ k) = ⊤, - from ih (finrank K ER) h_dim_ER f' rfl, - -- The induction hypothesis gives us a statement about subspaces of `ER`. We can transfer this - -- to a statement about subspaces of `V` via `submodule.subtype`: - have ih_ER' : (⨆ (μ : K) (k : ℕ), (f'.generalized_eigenspace μ k).map ER.subtype) = ER, - by simp only [(submodule.map_supr _ _).symm, ih_ER, submodule.map_subtype_top ER], - -- Moreover, every generalized eigenspace of `f'` is contained in the corresponding generalized - -- eigenspace of `f`. - have hff' : ∀ μ k, - (f'.generalized_eigenspace μ k).map ER.subtype ≤ f.generalized_eigenspace μ k, - { intros, - rw generalized_eigenspace_restrict, - apply submodule.map_comap_le }, - -- It follows that `ER` is contained in the span of all generalized eigenvectors. - have hER : ER ≤ ⨆ (μ : K) (k : ℕ), f.generalized_eigenspace μ k, - { rw ← ih_ER', - exact supr₂_mono hff' }, - -- `ES` is contained in this span by definition. - have hES : ES ≤ ⨆ (μ : K) (k : ℕ), f.generalized_eigenspace μ k, - from le_trans - (le_supr (λ k, f.generalized_eigenspace μ₀ k) (finrank K V)) - (le_supr (λ (μ : K), ⨆ (k : ℕ), f.generalized_eigenspace μ k) μ₀), - -- Moreover, we know that `ER` and `ES` are disjoint. - have h_disjoint : disjoint ER ES, - from generalized_eigenvec_disjoint_range_ker f μ₀, - -- Since the dimensions of `ER` and `ES` add up to the dimension of `V`, it follows that the - -- span of all generalized eigenvectors is all of `V`. - show (⨆ (μ : K) (k : ℕ), f.generalized_eigenspace μ k) = ⊤, - { rw [←top_le_iff, ←submodule.eq_top_of_disjoint ER ES h_dim_add h_disjoint], - apply sup_le hER hES } } -end - end End end module diff --git a/src/linear_algebra/eigenspace/is_alg_closed.lean b/src/linear_algebra/eigenspace/is_alg_closed.lean new file mode 100644 index 0000000000000..d8fb10800d9b6 --- /dev/null +++ b/src/linear_algebra/eigenspace/is_alg_closed.lean @@ -0,0 +1,115 @@ +/- +Copyright (c) 2020 Alexander Bentkamp. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Alexander Bentkamp +-/ + +import linear_algebra.eigenspace.basic +import field_theory.is_alg_closed.spectrum + +/-! +# Eigenvectors and eigenvalues over algebraically closed fields. + +* Every linear operator on a vector space over an algebraically closed field has an eigenvalue. +* The generalized eigenvectors span the entire vector space. + +## References + +* [Sheldon Axler, *Linear Algebra Done Right*][axler2015] +* https://en.wikipedia.org/wiki/Eigenvalues_and_eigenvectors + +## Tags + +eigenspace, eigenvector, eigenvalue, eigen +-/ + +universes u v w + +namespace module +namespace End + +open finite_dimensional + +variables {K : Type v} {V : Type w} [field K] [add_comm_group V] [module K V] + +/-- Every linear operator on a vector space over an algebraically closed field has + an eigenvalue. -/ +-- This is Lemma 5.21 of [axler2015], although we are no longer following that proof. +lemma exists_eigenvalue [is_alg_closed K] [finite_dimensional K V] [nontrivial V] (f : End K V) : + ∃ (c : K), f.has_eigenvalue c := +by { simp_rw has_eigenvalue_iff_mem_spectrum, + exact spectrum.nonempty_of_is_alg_closed_of_finite_dimensional K f } + +noncomputable instance [is_alg_closed K] [finite_dimensional K V] [nontrivial V] (f : End K V) : + inhabited f.eigenvalues := +⟨⟨f.exists_eigenvalue.some, f.exists_eigenvalue.some_spec⟩⟩ + +/-- The generalized eigenvectors span the entire vector space (Lemma 8.21 of [axler2015]). -/ +lemma supr_generalized_eigenspace_eq_top [is_alg_closed K] [finite_dimensional K V] (f : End K V) : + (⨆ (μ : K) (k : ℕ), f.generalized_eigenspace μ k) = ⊤ := +begin + -- We prove the claim by strong induction on the dimension of the vector space. + unfreezingI { induction h_dim : finrank K V using nat.strong_induction_on + with n ih generalizing V }, + cases n, + -- If the vector space is 0-dimensional, the result is trivial. + { rw ←top_le_iff, + simp only [finrank_eq_zero.1 (eq.trans (finrank_top _ _) h_dim), bot_le] }, + -- Otherwise the vector space is nontrivial. + { haveI : nontrivial V := finrank_pos_iff.1 (by { rw h_dim, apply nat.zero_lt_succ }), + -- Hence, `f` has an eigenvalue `μ₀`. + obtain ⟨μ₀, hμ₀⟩ : ∃ μ₀, f.has_eigenvalue μ₀ := exists_eigenvalue f, + -- We define `ES` to be the generalized eigenspace + let ES := f.generalized_eigenspace μ₀ (finrank K V), + -- and `ER` to be the generalized eigenrange. + let ER := f.generalized_eigenrange μ₀ (finrank K V), + -- `f` maps `ER` into itself. + have h_f_ER : ∀ (x : V), x ∈ ER → f x ∈ ER, + from λ x hx, map_generalized_eigenrange_le (submodule.mem_map_of_mem hx), + -- Therefore, we can define the restriction `f'` of `f` to `ER`. + let f' : End K ER := f.restrict h_f_ER, + -- The dimension of `ES` is positive + have h_dim_ES_pos : 0 < finrank K ES, + { dsimp only [ES], + rw h_dim, + apply pos_finrank_generalized_eigenspace_of_has_eigenvalue hμ₀ (nat.zero_lt_succ n) }, + -- and the dimensions of `ES` and `ER` add up to `finrank K V`. + have h_dim_add : finrank K ER + finrank K ES = finrank K V, + { apply linear_map.finrank_range_add_finrank_ker }, + -- Therefore the dimension `ER` mus be smaller than `finrank K V`. + have h_dim_ER : finrank K ER < n.succ, by linarith, + -- This allows us to apply the induction hypothesis on `ER`: + have ih_ER : (⨆ (μ : K) (k : ℕ), f'.generalized_eigenspace μ k) = ⊤, + from ih (finrank K ER) h_dim_ER f' rfl, + -- The induction hypothesis gives us a statement about subspaces of `ER`. We can transfer this + -- to a statement about subspaces of `V` via `submodule.subtype`: + have ih_ER' : (⨆ (μ : K) (k : ℕ), (f'.generalized_eigenspace μ k).map ER.subtype) = ER, + by simp only [(submodule.map_supr _ _).symm, ih_ER, submodule.map_subtype_top ER], + -- Moreover, every generalized eigenspace of `f'` is contained in the corresponding generalized + -- eigenspace of `f`. + have hff' : ∀ μ k, + (f'.generalized_eigenspace μ k).map ER.subtype ≤ f.generalized_eigenspace μ k, + { intros, + rw generalized_eigenspace_restrict, + apply submodule.map_comap_le }, + -- It follows that `ER` is contained in the span of all generalized eigenvectors. + have hER : ER ≤ ⨆ (μ : K) (k : ℕ), f.generalized_eigenspace μ k, + { rw ← ih_ER', + exact supr₂_mono hff' }, + -- `ES` is contained in this span by definition. + have hES : ES ≤ ⨆ (μ : K) (k : ℕ), f.generalized_eigenspace μ k, + from le_trans + (le_supr (λ k, f.generalized_eigenspace μ₀ k) (finrank K V)) + (le_supr (λ (μ : K), ⨆ (k : ℕ), f.generalized_eigenspace μ k) μ₀), + -- Moreover, we know that `ER` and `ES` are disjoint. + have h_disjoint : disjoint ER ES, + from generalized_eigenvec_disjoint_range_ker f μ₀, + -- Since the dimensions of `ER` and `ES` add up to the dimension of `V`, it follows that the + -- span of all generalized eigenvectors is all of `V`. + show (⨆ (μ : K) (k : ℕ), f.generalized_eigenspace μ k) = ⊤, + { rw [←top_le_iff, ←submodule.eq_top_of_disjoint ER ES h_dim_add h_disjoint], + apply sup_le hER hES } } +end + +end End +end module diff --git a/src/linear_algebra/eigenspace/minpoly.lean b/src/linear_algebra/eigenspace/minpoly.lean new file mode 100644 index 0000000000000..feafdd2ee815f --- /dev/null +++ b/src/linear_algebra/eigenspace/minpoly.lean @@ -0,0 +1,113 @@ +/- +Copyright (c) 2020 Alexander Bentkamp. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Alexander Bentkamp +-/ + +import linear_algebra.eigenspace.basic +import field_theory.minpoly.field + +/-! +# Eigenvalues are the roots of the minimal polynomial. + +## Tags + +eigenvalue, minimal polynomial +-/ + +universes u v w + +namespace module +namespace End + +open polynomial finite_dimensional +open_locale polynomial + +variables {K : Type v} {V : Type w} [field K] [add_comm_group V] [module K V] + +lemma eigenspace_aeval_polynomial_degree_1 + (f : End K V) (q : K[X]) (hq : degree q = 1) : + eigenspace f (- q.coeff 0 / q.leading_coeff) = (aeval f q).ker := +calc + eigenspace f (- q.coeff 0 / q.leading_coeff) + = (q.leading_coeff • f - algebra_map K (End K V) (- q.coeff 0)).ker + : by { rw eigenspace_div, intro h, rw leading_coeff_eq_zero_iff_deg_eq_bot.1 h at hq, cases hq } + ... = (aeval f (C q.leading_coeff * X + C (q.coeff 0))).ker + : by { rw [C_mul', aeval_def], simp [algebra_map, algebra.to_ring_hom], } + ... = (aeval f q).ker + : by rwa ← eq_X_add_C_of_degree_eq_one + +lemma ker_aeval_ring_hom'_unit_polynomial + (f : End K V) (c : (K[X])ˣ) : + (aeval f (c : K[X])).ker = ⊥ := +begin + rw polynomial.eq_C_of_degree_eq_zero (degree_coe_units c), + simp only [aeval_def, eval₂_C], + apply ker_algebra_map_End, + apply coeff_coe_units_zero_ne_zero c +end + +theorem aeval_apply_of_has_eigenvector {f : End K V} + {p : K[X]} {μ : K} {x : V} (h : f.has_eigenvector μ x) : + aeval f p x = (p.eval μ) • x := +begin + apply p.induction_on, + { intro a, simp [module.algebra_map_End_apply] }, + { intros p q hp hq, simp [hp, hq, add_smul] }, + { intros n a hna, + rw [mul_comm, pow_succ, mul_assoc, alg_hom.map_mul, linear_map.mul_apply, mul_comm, hna], + simp only [mem_eigenspace_iff.1 h.1, smul_smul, aeval_X, eval_mul, eval_C, eval_pow, eval_X, + linear_map.map_smulₛₗ, ring_hom.id_apply, mul_comm] } +end + +theorem is_root_of_has_eigenvalue {f : End K V} {μ : K} (h : f.has_eigenvalue μ) : + (minpoly K f).is_root μ := +begin + rcases (submodule.ne_bot_iff _).1 h with ⟨w, ⟨H, ne0⟩⟩, + refine or.resolve_right (smul_eq_zero.1 _) ne0, + simp [← aeval_apply_of_has_eigenvector ⟨H, ne0⟩, minpoly.aeval K f], +end + +variables [finite_dimensional K V] (f : End K V) + +variables {f} {μ : K} + +theorem has_eigenvalue_of_is_root (h : (minpoly K f).is_root μ) : + f.has_eigenvalue μ := +begin + cases dvd_iff_is_root.2 h with p hp, + rw [has_eigenvalue, eigenspace], + intro con, + cases (linear_map.is_unit_iff_ker_eq_bot _).2 con with u hu, + have p_ne_0 : p ≠ 0, + { intro con, + apply minpoly.ne_zero f.is_integral, + rw [hp, con, mul_zero] }, + have h_deg := minpoly.degree_le_of_ne_zero K f p_ne_0 _, + { rw [hp, degree_mul, degree_X_sub_C, polynomial.degree_eq_nat_degree p_ne_0] at h_deg, + norm_cast at h_deg, + linarith, }, + { have h_aeval := minpoly.aeval K f, + revert h_aeval, + simp [hp, ← hu] }, +end + +theorem has_eigenvalue_iff_is_root : + f.has_eigenvalue μ ↔ (minpoly K f).is_root μ := +⟨is_root_of_has_eigenvalue, has_eigenvalue_of_is_root⟩ + +/-- An endomorphism of a finite-dimensional vector space has finitely many eigenvalues. -/ +noncomputable instance (f : End K V) : fintype f.eigenvalues := +set.finite.fintype +begin + have h : minpoly K f ≠ 0 := minpoly.ne_zero f.is_integral, + convert (minpoly K f).root_set_finite K, + ext μ, + have : (μ ∈ {μ : K | f.eigenspace μ = ⊥ → false}) ↔ ¬f.eigenspace μ = ⊥ := by tauto, + convert rfl.mpr this, + simp [polynomial.root_set_def, polynomial.mem_roots h, ← has_eigenvalue_iff_is_root, + has_eigenvalue] +end + +end End +end module From 81843c08659063f91486be95a617b5e9ec93c9da Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Wed, 7 Jun 2023 07:55:03 +0000 Subject: [PATCH 1152/1291] feat (number_theory/zeta_function): func eqn + link to Bernoulli nos (#19158) Functional equation relating `zeta (1 - s)` to `zeta s`, and formula for `zeta (-k)` for nat `k` in terms of Bernoulli numbers. --- src/number_theory/zeta_function.lean | 235 ++++++++++++++++++++++++++- 1 file changed, 233 insertions(+), 2 deletions(-) diff --git a/src/number_theory/zeta_function.lean b/src/number_theory/zeta_function.lean index 2a1c9d2101337..4254a814864f1 100644 --- a/src/number_theory/zeta_function.lean +++ b/src/number_theory/zeta_function.lean @@ -5,6 +5,7 @@ Authors: David Loeffler -/ import analysis.special_functions.gamma.beta import number_theory.modular_forms.jacobi_theta +import number_theory.zeta_values /-! # Definition of the Riemann zeta function @@ -29,6 +30,12 @@ I haven't checked exactly what they are). * `differentiable_at_riemann_zeta` : the function `ζ(s)` is differentiable away from `s = 1`. * `zeta_eq_tsum_of_one_lt_re` : for `1 < re s`, we have `ζ(s) = ∑' (n : ℕ), 1 / (n + 1) ^ s`. +* `riemann_completed_zeta₀_one_sub`, `riemann_completed_zeta_one_sub`, and `riemann_zeta_one_sub` : + functional equation relating values at `s` and `1 - s` +* `riemann_zeta_neg_nat_eq_bernoulli` : for any `k ∈ ℕ` we have the formula + `riemann_zeta (-k) = (-1) ^ k * bernoulli (k + 1) / (k + 1)` +* `riemann_zeta_two_mul_nat`: formula for `ζ(2 * k)` for `k ∈ ℕ, k ≠ 0` in terms of Bernoulli + numbers ## Outline of proofs: @@ -38,12 +45,18 @@ completed zeta function. The second is obtained by subtracting a linear combinat the interval `Ioc 0 1` to give a function with exponential decay at both `0` and `∞`. We then define `riemann_completed_zeta₀` as the Mellin transform of the second zeta kernel, and define `riemann_completed_zeta` and `riemann_zeta` from this. + +Since `zeta_kernel₂` has rapid decay and satisfies a functional equation relating its values at `t` +and `1 / t`, we deduce the analyticity of `riemann_completed_zeta₀` and the functional equation +relating its values at `s` and `1 - s`. On the other hand, since `zeta_kernel₁` can be expanded in +powers of `exp (-π * t)` and the Mellin transform integrated term-by-term, we obtain the relation +to the naive Dirichlet series `∑' (n : ℕ), 1 / (n + 1) ^ s`. -/ open measure_theory set filter asymptotics topological_space real asymptotics open complex (hiding exp norm_eq_abs abs_of_nonneg abs_two continuous_exp) -open_locale topology real +open_locale topology real nat noncomputable theory @@ -522,7 +535,7 @@ end /-- The Riemann zeta function agrees with the naive Dirichlet-series definition when the latter converges. (Note that this is false without the assumption: when `re s ≤ 1` the sum is divergent, and we use a different definition to obtain the analytic continuation to all `s`.) -/ -theorem zeta_eq_tsum_of_one_lt_re {s : ℂ} (hs : 1 < re s) : +theorem zeta_eq_tsum_one_div_nat_add_one_cpow {s : ℂ} (hs : 1 < re s) : riemann_zeta s = ∑' (n : ℕ), 1 / (n + 1) ^ s := begin have : s ≠ 0, by { contrapose! hs, rw [hs, zero_re], exact zero_le_one }, @@ -534,3 +547,221 @@ begin { rw [ne.def, cpow_eq_zero_iff, not_and_distrib, ←ne.def, of_real_ne_zero], exact or.inl (pi_pos.ne') } end + +/-- Alternate formulation of `zeta_eq_tsum_one_div_nat_add_one_cpow` without the `+ 1`, using the +fact that for `s ≠ 0` we define `0 ^ s = 0`. -/ +lemma zeta_eq_tsum_one_div_nat_cpow {s : ℂ} (hs : 1 < re s) : + riemann_zeta s = ∑' (n : ℕ), 1 / n ^ s := +begin + have hs' : s ≠ 0, by { contrapose! hs, rw [hs, zero_re], exact zero_le_one }, + rw [tsum_eq_zero_add], + { simp_rw [nat.cast_zero, zero_cpow hs', div_zero, zero_add, + zeta_eq_tsum_one_div_nat_add_one_cpow hs, nat.cast_add, nat.cast_one] }, + { rw ←summable_norm_iff, + simp_rw [norm_div, norm_one, complex.norm_eq_abs, ←of_real_nat_cast, + abs_cpow_eq_rpow_re_of_nonneg (nat.cast_nonneg _) (zero_lt_one.trans hs).ne', + summable_one_div_nat_rpow], + assumption } +end + +/-- Special case of `zeta_eq_tsum_one_div_nat_cpow` when the argument is in `ℕ`, so the power +function can be expressed using naïve `pow` rather than `cpow`. -/ +lemma zeta_nat_eq_tsum_of_gt_one {k : ℕ} (hk : 1 < k) : riemann_zeta k = ∑' (n : ℕ), 1 / n ^ k := +by simp only [zeta_eq_tsum_one_div_nat_cpow (by rwa [←of_real_nat_cast, of_real_re, ←nat.cast_one, + nat.cast_lt] : 1 < re k), cpow_nat_cast] + +/-- Explicit formula for `ζ (2 * k)`, for `k ∈ ℕ` with `k ≠ 0`: we have +`ζ (2 * k) = (-1) ^ (k + 1) * 2 ^ (2 * k - 1) * π ^ (2 * k) * bernoulli (2 * k) / (2 * k)!`. +Compare `has_sum_zeta_nat` for a version formulated explicitly as a sum, and +`riemann_zeta_neg_nat_eq_bernoulli` for values at negative integers (equivalent to the above via +the functional equation). -/ +lemma riemann_zeta_two_mul_nat {k : ℕ} (hk : k ≠ 0) : + riemann_zeta (2 * k) = + (-1) ^ (k + 1) * 2 ^ (2 * k - 1) * π ^ (2 * k) * bernoulli (2 * k) / (2 * k)! := +begin + convert congr_arg (coe : ℝ → ℂ) (has_sum_zeta_nat hk).tsum_eq, + { rw [←nat.cast_two, ←nat.cast_mul, zeta_nat_eq_tsum_of_gt_one], + { push_cast }, + { refine (one_lt_two).trans_le _, + conv_lhs { rw ←mul_one 2 }, + rwa [mul_le_mul_left (zero_lt_two' ℕ), nat.one_le_iff_ne_zero] } }, + { push_cast } +end + +lemma riemann_zeta_two : riemann_zeta 2 = π ^ 2 / 6 := +begin + convert congr_arg coe has_sum_zeta_two.tsum_eq, + { rw [←nat.cast_two, zeta_nat_eq_tsum_of_gt_one one_lt_two, of_real_tsum], + push_cast }, + { push_cast } +end + +lemma riemann_zeta_four : riemann_zeta 4 = π ^ 4 / 90 := +begin + convert congr_arg coe has_sum_zeta_four.tsum_eq, + { rw [←nat.cast_one, ←nat.cast_bit0, ←nat.cast_bit0, zeta_nat_eq_tsum_of_gt_one + (by norm_num : 1 < 4), of_real_tsum], + push_cast }, + { push_cast } +end + +/-! +## Functional equation +-/ + +/-- Riemann zeta functional equation, formulated for `Λ₀`: for any complex `s` we have +`Λ₀(1 - s) = Λ₀ s`. -/ +lemma riemann_completed_zeta₀_one_sub (s : ℂ) : + riemann_completed_zeta₀ (1 - s) = riemann_completed_zeta₀ s := +begin + have := mellin_comp_rpow (zeta_kernel₂) (s / 2 - 1 / 2) neg_one_lt_zero.ne, + simp_rw [rpow_neg_one, ←one_div, abs_neg, abs_one, div_one, one_smul, of_real_neg, + of_real_one, div_neg, div_one, neg_sub] at this, + conv_lhs { rw [riemann_completed_zeta₀, sub_div, ←this] }, + refine set_integral_congr measurable_set_Ioi (λ t ht, _), + simp_rw [zeta_kernel₂_one_div ht, smul_eq_mul, ←mul_assoc, sqrt_eq_rpow, + of_real_cpow (le_of_lt ht), ←cpow_add _ _ (of_real_ne_zero.mpr $ ne_of_gt ht)], + congr' 2, + push_cast, + ring, +end + +/-- Riemann zeta functional equation, formulated for `Λ`: for any complex `s` we have +`Λ (1 - s) = Λ s`. -/ +lemma riemann_completed_zeta_one_sub (s : ℂ) : + riemann_completed_zeta (1 - s) = riemann_completed_zeta s := +by simp_rw [riemann_completed_zeta, riemann_completed_zeta₀_one_sub, sub_add, + (by abel : 1 - s - 1 = -s), (by abel : 1 - s = -(s - 1)), div_neg, neg_sub_neg] + +/-- Riemann zeta functional equation, formulated for `ζ`: if `1 - s ∉ ℕ`, then we have +`ζ (1 - s) = 2 ^ (1 - s) * π ^ (-s) * Γ s * sin (π * (1 - s) / 2) * ζ s`. -/ +lemma riemann_zeta_one_sub {s : ℂ} (hs : ∀ (n : ℕ), s ≠ -n) (hs' : s ≠ 1) : + riemann_zeta (1 - s) = + 2 ^ (1 - s) * π ^ (-s) * Gamma s * sin (π * (1 - s) / 2) * riemann_zeta s := +begin + -- Deducing this from the previous formulations is quite involved. The proof uses two + -- nontrivial facts (the doubling formula and reflection formula for Gamma) and a lot of careful + -- rearrangement, requiring several non-vanishing statements as input to `field_simp`. + have hs_ne : s ≠ 0, by { contrapose! hs, rw hs, exact ⟨0, by rw [nat.cast_zero, neg_zero]⟩ }, + have h_sqrt : (sqrt π : ℂ) ≠ 0, from of_real_ne_zero.mpr (sqrt_ne_zero'.mpr pi_pos), + have h_pow : (2 : ℂ) ^ (s - 1) ≠ 0, + { rw [ne.def, cpow_eq_zero_iff, not_and_distrib], exact or.inl two_ne_zero }, + have h_Ga_ne1 : Gamma (s / 2) ≠ 0, + { rw [ne.def, complex.Gamma_eq_zero_iff], + contrapose! hs, + obtain ⟨m, hm⟩ := hs, + rw [div_eq_iff (two_ne_zero' ℂ), ←nat.cast_two, neg_mul, ←nat.cast_mul] at hm, + exact ⟨m * 2, by rw hm⟩ }, + have h_Ga_eq : Gamma s = Gamma (s / 2) * Gamma ((s + 1) / 2) * 2 ^ (s - 1) / sqrt π, + { rw [add_div, complex.Gamma_mul_Gamma_add_half, mul_div_cancel' _ (two_ne_zero' ℂ), + (by ring : 1 - s = -(s - 1)), cpow_neg, ←div_eq_mul_inv, eq_div_iff h_sqrt, + div_mul_eq_mul_div₀, div_mul_cancel _ h_pow] }, + have h_Ga_ne3 : Gamma ((s + 1) / 2) ≠ 0, + { have h_Ga_aux : Gamma s ≠ 0, from complex.Gamma_ne_zero hs, + contrapose! h_Ga_aux, + rw [h_Ga_eq, h_Ga_aux, mul_zero, zero_mul, zero_div] }, + rw [riemann_zeta, function.update_noteq (by rwa [sub_ne_zero, ne_comm] : 1 - s ≠ 0), + function.update_noteq hs_ne, riemann_completed_zeta_one_sub, mul_div, eq_div_iff h_Ga_ne1, + mul_comm, ←mul_div_assoc], + -- Now rule out case of s = positive odd integer & deduce further non-vanishing statements + by_cases hs_pos_odd : ∃ (n : ℕ), s = 1 + 2 * n, + { -- Note the case n = 0 (i.e. s = 1) works OK here, but only because we have used + -- `function.update_noteq` to change the goal; the original goal is genuinely false for s = 1. + obtain ⟨n, rfl⟩ := hs_pos_odd, + have : (1 - (1 + 2 * (n : ℂ))) / 2 = -↑n, + { rw [←sub_sub, sub_self, zero_sub, neg_div, mul_div_cancel_left _ (two_ne_zero' ℂ)] }, + rw [this, complex.Gamma_neg_nat_eq_zero, div_zero], + have : (π : ℂ) * (1 - (1 + 2 * ↑n)) / 2 = ↑(-n : ℤ) * π, + { push_cast, field_simp, ring }, + rw [this, complex.sin_int_mul_pi, mul_zero, zero_mul] }, + have h_Ga_ne4 : Gamma ((1 - s) / 2) ≠ 0, + { rw [ne.def, complex.Gamma_eq_zero_iff], + contrapose! hs_pos_odd, + obtain ⟨m, hm⟩ := hs_pos_odd, + rw [div_eq_iff (two_ne_zero' ℂ), sub_eq_iff_eq_add, neg_mul, ←sub_eq_neg_add, + eq_sub_iff_add_eq] at hm, + exact ⟨m, by rw [←hm, mul_comm]⟩ }, + -- At last the main proof + rw show sin (↑π * (1 - s) / 2) = π * (Gamma ((1 - s) / 2) * Gamma (s / 2 + 1 / 2))⁻¹, by + { have := congr_arg has_inv.inv (complex.Gamma_mul_Gamma_one_sub ((1 - s) / 2)).symm, + rwa [(by ring : 1 - (1 - s) / 2 = s / 2 + 1 / 2), inv_div, + div_eq_iff (of_real_ne_zero.mpr pi_pos.ne'), mul_comm _ ↑π, mul_div_assoc'] at this }, + rw [(by rw ←neg_sub : (2 : ℂ) ^ (1 - s) = 2 ^ -(s - 1)), cpow_neg, h_Ga_eq], + suffices : (π : ℂ) ^ ((1 - s) / 2) = π ^ -s * sqrt π * π ^ (s / 2), + { rw this, field_simp, ring_nf, rw [←of_real_pow, sq_sqrt pi_pos.le], ring }, + simp_rw [sqrt_eq_rpow, of_real_cpow pi_pos.le, ←cpow_add _ _ (of_real_ne_zero.mpr pi_pos.ne')], + congr' 1, + push_cast, + field_simp, + ring, +end + +lemma riemann_zeta_neg_nat_eq_bernoulli (k : ℕ) : + riemann_zeta (-k) = (-1) ^ k * bernoulli (k + 1) / (k + 1) := +begin + rcases nat.even_or_odd' k with ⟨m, rfl | rfl⟩, + { cases m, + { -- k = 0 : evaluate explicitly + rw [mul_zero, nat.cast_zero, pow_zero, one_mul, zero_add, neg_zero, zero_add, div_one, + bernoulli_one, riemann_zeta_zero, rat.cast_div, rat.cast_neg, rat.cast_one, + rat.cast_bit0, rat.cast_one] }, + { -- k = 2 * (m + 1) : both sides "trivially" zero + rw [nat.cast_mul, ←neg_mul, nat.cast_two, nat.cast_succ, + riemann_zeta_neg_two_mul_nat_add_one, bernoulli_eq_bernoulli'_of_ne_one], + swap, { apply ne_of_gt, norm_num }, + rw [bernoulli'_odd_eq_zero ⟨m + 1, rfl⟩ (by norm_num), rat.cast_zero, mul_zero, zero_div] } }, + { -- k = 2 * m + 1 : the interesting case + rw odd.neg_one_pow ⟨m, rfl⟩, + rw (show -(↑(2 * m + 1) : ℂ) = 1 - (2 * m + 2), by { push_cast, ring }), + rw riemann_zeta_one_sub, + rotate, + { intro n, + rw [(by norm_cast : (2 * (m : ℂ) + 2) = ↑(2 * m + 2)), ←int.cast_neg_nat_cast, + ←int.cast_coe_nat, ne.def, int.cast_inj], + apply ne_of_gt, + refine lt_of_le_of_lt (by norm_num : (-n : ℤ) ≤ 0) (by positivity) }, + { rw [(by norm_cast : (2 * (m : ℂ) + 2) = ↑(2 * m + 2)), ne.def, nat.cast_eq_one], norm_num }, + -- get rid of sine term + rw show complex.sin (↑π * (1 - (2 * ↑m + 2)) / 2) = -(-1) ^ m, + { rw (by { field_simp, ring } : (π : ℂ) * (1 - (2 * ↑m + 2)) / 2 = π / 2 - (π * m + π)), + rw [complex.sin_pi_div_two_sub, complex.cos_add_pi, neg_inj], + rcases nat.even_or_odd' m with ⟨t, rfl | rfl⟩, + { rw [pow_mul, neg_one_sq, one_pow], + convert complex.cos_nat_mul_two_pi t using 2, push_cast, ring }, + { rw [pow_add, pow_one, pow_mul, neg_one_sq, one_pow, one_mul], + convert complex.cos_nat_mul_two_pi_add_pi t using 2, push_cast, ring } }, + -- substitute in what we know about zeta values at positive integers + have step1 := congr_arg (coe : ℝ → ℂ) (has_sum_zeta_nat (by norm_num : m + 1 ≠ 0)).tsum_eq, + have step2 := zeta_nat_eq_tsum_of_gt_one (by { rw mul_add, norm_num } : 1 < 2 * (m + 1)), + simp_rw [of_real_tsum, of_real_div, of_real_one, of_real_pow, of_real_nat_cast] at step1, + rw [step1, (by norm_cast : (↑(2 * (m + 1)) : ℂ) = 2 * ↑m + 2)] at step2, + rw [step2, mul_div], + -- now the rest is just a lengthy but elementary rearrangement + rw show ((2 * (m + 1))! : ℂ) = Gamma (2 * m + 2) * (↑(2 * m + 1) + 1), by + { rw [(by { push_cast, ring } : (2 * m + 2 : ℂ) = ↑(2 * m + 1) + 1), + complex.Gamma_nat_eq_factorial, + (by ring : 2 * (m + 1) = (2 * m + 1) + 1), nat.factorial_succ, nat.cast_mul, mul_comm], + push_cast }, + rw [←div_div, neg_one_mul], + congr' 1, + rw [div_eq_iff (Gamma_ne_zero_of_re_pos _)], + swap, { rw [(by push_cast : 2 * (m : ℂ) + 2 = ↑(2 * (m : ℝ) + 2)), of_real_re], positivity }, + simp_rw [of_real_mul, ←mul_assoc, of_real_rat_cast, mul_add, nat.add_assoc, mul_one, + one_add_one_eq_two, mul_neg, neg_mul, neg_inj], + conv_rhs { rw mul_comm }, + congr' 1, + rw [of_real_pow, of_real_neg, of_real_one, pow_add, neg_one_sq, mul_one], + conv_lhs { congr, congr, + rw [mul_assoc, ←pow_add, ←two_mul, pow_mul, neg_one_sq, one_pow, mul_one] }, + rw show (2 : ℂ) ^ (1 - (2 * (m : ℂ) + 2)) = (↑((2 : ℝ) ^ ((2 * m + 2) - 1)))⁻¹, + { rw [of_real_pow, ←cpow_nat_cast, ←cpow_neg, of_real_bit0, of_real_one], + congr' 1, + rw [nat.add_sub_assoc one_le_two, nat.cast_add, nat.cast_mul, nat.cast_two, + (by norm_num : 2 - 1 = 1)], + push_cast, ring }, + rw show (π : ℂ) ^ -(2 * (m : ℂ) + 2) = (↑(π ^ (2 * m + 2)))⁻¹, + { rw [of_real_pow, ←cpow_nat_cast, ←cpow_neg, nat.cast_add, nat.cast_mul, nat.cast_two] }, + rw (by { intros, ring } : ∀ (a b c d e : ℂ), a * b * c * d * e = (a * d) * (b * e) * c), + rw [inv_mul_cancel (of_real_ne_zero.mpr $ pow_ne_zero _ pi_pos.ne'), + inv_mul_cancel (of_real_ne_zero.mpr $ pow_ne_zero _ two_ne_zero), one_mul, one_mul] } +end From 0dc4079202c28226b2841a51eb6d3cc2135bb80f Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Wed, 7 Jun 2023 10:28:17 +0000 Subject: [PATCH 1153/1291] refactor(geometry/manifold/cont_mdiff_map): refactor to reduce imports (#19164) See https://tqft.net/mathlib4/2023-06-07/geometry.manifold.algebra.left_invariant_derivation.pdf for the motivation. I'm happy if this one is shot down. :-) Co-authored-by: Scott Morrison --- src/geometry/manifold/cont_mdiff_map.lean | 14 +---------- src/geometry/manifold/cont_mdiff_mfderiv.lean | 24 +++++++++++++++++++ src/geometry/manifold/diffeomorph.lean | 1 + src/geometry/manifold/instances/sphere.lean | 1 + 4 files changed, 27 insertions(+), 13 deletions(-) diff --git a/src/geometry/manifold/cont_mdiff_map.lean b/src/geometry/manifold/cont_mdiff_map.lean index 05bad7919e598..b064885bcbda5 100644 --- a/src/geometry/manifold/cont_mdiff_map.lean +++ b/src/geometry/manifold/cont_mdiff_map.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Nicolò Cavalleri -/ -import geometry.manifold.cont_mdiff_mfderiv +import geometry.manifold.cont_mdiff import topology.continuous_function.basic /-! @@ -66,18 +66,6 @@ protected lemma cont_mdiff (f : C^n⟮I, M; I', M'⟯) : protected lemma smooth (f : C^∞⟮I, M; I', M'⟯) : smooth I I' f := f.cont_mdiff_to_fun -protected lemma mdifferentiable' (f : C^n⟮I, M; I', M'⟯) (hn : 1 ≤ n) : - mdifferentiable I I' f := -f.cont_mdiff.mdifferentiable hn - -protected lemma mdifferentiable (f : C^∞⟮I, M; I', M'⟯) : - mdifferentiable I I' f := -f.cont_mdiff.mdifferentiable le_top - -protected lemma mdifferentiable_at (f : C^∞⟮I, M; I', M'⟯) {x} : - mdifferentiable_at I I' f x := -f.mdifferentiable x - lemma coe_inj ⦃f g : C^n⟮I, M; I', M'⟯⦄ (h : (f : M → M') = g) : f = g := by cases f; cases g; cases h; refl diff --git a/src/geometry/manifold/cont_mdiff_mfderiv.lean b/src/geometry/manifold/cont_mdiff_mfderiv.lean index 343ec2c230329..3ba92b1b57c6d 100644 --- a/src/geometry/manifold/cont_mdiff_mfderiv.lean +++ b/src/geometry/manifold/cont_mdiff_mfderiv.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel, Floris van Doorn -/ import geometry.manifold.mfderiv +import geometry.manifold.cont_mdiff_map /-! ### Interactions between differentiability, smoothness and manifold derivatives @@ -609,3 +610,26 @@ begin end end tangent_bundle + +namespace cont_mdiff_map + +-- These helpers for dot notation have been moved here from `geometry.manifold.cont_mdiff_map` +-- to avoid needing to import `geometry.manifold.cont_mdiff_mfderiv` there. +-- (However as a consequence we import `geometry.manifold.cont_mdiff_map` here now.) +-- They could be moved to another file (perhaps a new file) if desired. + +open_locale manifold + +protected lemma mdifferentiable' (f : C^n⟮I, M; I', M'⟯) (hn : 1 ≤ n) : + mdifferentiable I I' f := +f.cont_mdiff.mdifferentiable hn + +protected lemma mdifferentiable (f : C^∞⟮I, M; I', M'⟯) : + mdifferentiable I I' f := +f.cont_mdiff.mdifferentiable le_top + +protected lemma mdifferentiable_at (f : C^∞⟮I, M; I', M'⟯) {x} : + mdifferentiable_at I I' f x := +f.mdifferentiable x + +end cont_mdiff_map diff --git a/src/geometry/manifold/diffeomorph.lean b/src/geometry/manifold/diffeomorph.lean index a02d3e5dc4d45..49a538a391b4a 100644 --- a/src/geometry/manifold/diffeomorph.lean +++ b/src/geometry/manifold/diffeomorph.lean @@ -5,6 +5,7 @@ Authors: Nicolò Cavalleri, Yury Kudryashov -/ import geometry.manifold.cont_mdiff_map +import geometry.manifold.cont_mdiff_mfderiv /-! # Diffeomorphisms diff --git a/src/geometry/manifold/instances/sphere.lean b/src/geometry/manifold/instances/sphere.lean index 23bae7aa168e2..e274632af5e28 100644 --- a/src/geometry/manifold/instances/sphere.lean +++ b/src/geometry/manifold/instances/sphere.lean @@ -10,6 +10,7 @@ import analysis.inner_product_space.calculus import analysis.inner_product_space.pi_L2 import geometry.manifold.algebra.lie_group import geometry.manifold.instances.real +import geometry.manifold.cont_mdiff_mfderiv /-! # Manifold structure on the sphere From 3b92d54a05ee592aa2c6181a4e76b1bb7cc45d0b Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Wed, 7 Jun 2023 11:35:18 +0000 Subject: [PATCH 1154/1291] chore(probability/kernel/composition): redefine kernel.comp using measure.bind, remove kernel.map_measure (#19160) We replace the definition of `kernel.comp` with a new one using `measure.bind`: it removes the need for `is_s_finite_kernel` hypotheses in the definition and most lemmas. When the kernels are s-finite, the new definition coincides with the old one. We remove `kernel.map_measure` because it is exactly the same as `measure.bind` applied to a kernel. Co-authored-by: RemyDegenne --- src/probability/kernel/composition.lean | 51 ++++++----- src/probability/kernel/invariance.lean | 115 ++++++------------------ 2 files changed, 56 insertions(+), 110 deletions(-) diff --git a/src/probability/kernel/composition.lean b/src/probability/kernel/composition.lean index ced5d02480471..17cc45818f356 100644 --- a/src/probability/kernel/composition.lean +++ b/src/probability/kernel/composition.lean @@ -13,8 +13,8 @@ We define * the composition-product `κ ⊗ₖ η` of two s-finite kernels `κ : kernel α β` and `η : kernel (α × β) γ`, a kernel from `α` to `β × γ`. * the map and comap of a kernel along a measurable function. -* the composition `η ∘ₖ κ` of s-finite kernels `κ : kernel α β` and `η : kernel β γ`, - a kernel from `α` to `γ`. +* the composition `η ∘ₖ κ` of kernels `κ : kernel α β` and `η : kernel β γ`, kernel from `α` to + `γ`. * the product `κ ×ₖ η` of s-finite kernels `κ : kernel α β` and `η : kernel α γ`, a kernel from `α` to `β × γ`. @@ -34,7 +34,7 @@ Kernels built from other kernels: `∫⁻ c, g c ∂(map κ f hf a) = ∫⁻ b, g (f b) ∂(κ a)` * `comap (κ : kernel α β) (f : γ → α) (hf : measurable f) : kernel γ β` `∫⁻ b, g b ∂(comap κ f hf c) = ∫⁻ b, g b ∂(κ (f c))` -* `comp (η : kernel β γ) (κ : kernel α β) : kernel α γ`: composition of 2 s-finite kernels. +* `comp (η : kernel β γ) (κ : kernel α β) : kernel α γ`: composition of 2 kernels. We define a notation `η ∘ₖ κ = comp η κ`. `∫⁻ c, g c ∂((η ∘ₖ κ) a) = ∫⁻ b, ∫⁻ c, g c ∂(η b) ∂(κ a)` * `prod (κ : kernel α β) (η : kernel α γ) : kernel α (β × γ)`: product of 2 s-finite kernels. @@ -741,69 +741,72 @@ include mγ /-- Composition of two s-finite kernels. -/ noncomputable -def comp (η : kernel β γ) [is_s_finite_kernel η] (κ : kernel α β) [is_s_finite_kernel κ] : - kernel α γ := -snd (κ ⊗ₖ prod_mk_left α η) +def comp (η : kernel β γ) (κ : kernel α β) : kernel α γ := +{ val := λ a, (κ a).bind η, + property := (measure.measurable_bind' (kernel.measurable _)).comp (kernel.measurable _) } localized "infix (name := kernel.comp) ` ∘ₖ `:100 := probability_theory.kernel.comp" in probability_theory -lemma comp_apply (η : kernel β γ) [is_s_finite_kernel η] (κ : kernel α β) [is_s_finite_kernel κ] - (a : α) {s : set γ} (hs : measurable_set s) : +lemma comp_apply (η : kernel β γ) (κ : kernel α β) (a : α) : + (η ∘ₖ κ) a = (κ a).bind η := rfl + +lemma comp_apply' (η : kernel β γ) (κ : kernel α β) (a : α) {s : set γ} (hs : measurable_set s) : (η ∘ₖ κ) a s = ∫⁻ b, η b s ∂(κ a) := +by rw [comp_apply, measure.bind_apply hs (kernel.measurable _)] + +lemma comp_eq_snd_comp_prod (η : kernel β γ) [is_s_finite_kernel η] + (κ : kernel α β) [is_s_finite_kernel κ] : + η ∘ₖ κ = snd (κ ⊗ₖ prod_mk_left α η) := begin - rw [comp, snd_apply' _ _ hs, comp_prod_apply], + ext a s hs : 2, + rw [comp_apply' _ _ _ hs, snd_apply' _ _ hs, comp_prod_apply], swap, { exact measurable_snd hs, }, simp only [set.mem_set_of_eq, set.set_of_mem_eq, prod_mk_left_apply' _ _ s], end -lemma lintegral_comp (η : kernel β γ) [is_s_finite_kernel η] (κ : kernel α β) [is_s_finite_kernel κ] +lemma lintegral_comp (η : kernel β γ) (κ : kernel α β) (a : α) {g : γ → ℝ≥0∞} (hg : measurable g) : ∫⁻ c, g c ∂((η ∘ₖ κ) a) = ∫⁻ b, ∫⁻ c, g c ∂(η b) ∂(κ a) := -begin - rw [comp, lintegral_snd _ _ hg], - change ∫⁻ bc, (λ a b, g b) bc.fst bc.snd ∂((κ ⊗ₖ prod_mk_left α η) a) - = ∫⁻ b, ∫⁻ c, g c ∂(η b) ∂(κ a), - exact lintegral_comp_prod _ _ _ (hg.comp measurable_snd), -end +by rw [comp_apply, measure.lintegral_bind (kernel.measurable _) hg] instance is_markov_kernel.comp (η : kernel β γ) [is_markov_kernel η] (κ : kernel α β) [is_markov_kernel κ] : is_markov_kernel (η ∘ₖ κ) := -by { rw comp, apply_instance, } +by { rw comp_eq_snd_comp_prod, apply_instance, } instance is_finite_kernel.comp (η : kernel β γ) [is_finite_kernel η] (κ : kernel α β) [is_finite_kernel κ] : is_finite_kernel (η ∘ₖ κ) := -by { rw comp, apply_instance, } +by { rw comp_eq_snd_comp_prod, apply_instance, } instance is_s_finite_kernel.comp (η : kernel β γ) [is_s_finite_kernel η] (κ : kernel α β) [is_s_finite_kernel κ] : is_s_finite_kernel (η ∘ₖ κ) := -by { rw comp, apply_instance, } +by { rw comp_eq_snd_comp_prod, apply_instance, } /-- Composition of kernels is associative. -/ lemma comp_assoc {δ : Type*} {mδ : measurable_space δ} (ξ : kernel γ δ) [is_s_finite_kernel ξ] - (η : kernel β γ) [is_s_finite_kernel η] (κ : kernel α β) [is_s_finite_kernel κ] : + (η : kernel β γ) (κ : kernel α β) : (ξ ∘ₖ η ∘ₖ κ) = ξ ∘ₖ (η ∘ₖ κ) := begin refine ext_fun (λ a f hf, _), simp_rw [lintegral_comp _ _ _ hf, lintegral_comp _ _ _ hf.lintegral_kernel], end -lemma deterministic_comp_eq_map (hf : measurable f) (κ : kernel α β) [is_s_finite_kernel κ] : +lemma deterministic_comp_eq_map (hf : measurable f) (κ : kernel α β) : (deterministic f hf ∘ₖ κ) = map κ f hf := begin ext a s hs : 2, - simp_rw [map_apply' _ _ _ hs, comp_apply _ _ _ hs, deterministic_apply' hf _ hs, + simp_rw [map_apply' _ _ _ hs, comp_apply' _ _ _ hs, deterministic_apply' hf _ hs, lintegral_indicator_const_comp hf hs, one_mul], end -lemma comp_deterministic_eq_comap (κ : kernel α β) [is_s_finite_kernel κ] (hg : measurable g) : +lemma comp_deterministic_eq_comap (κ : kernel α β) (hg : measurable g) : (κ ∘ₖ deterministic g hg) = comap κ g hg := begin ext a s hs : 2, - simp_rw [comap_apply' _ _ _ s, comp_apply _ _ _ hs, deterministic_apply hg a, + simp_rw [comap_apply' _ _ _ s, comp_apply' _ _ _ hs, deterministic_apply hg a, lintegral_dirac' _ (kernel.measurable_coe κ hs)], end diff --git a/src/probability/kernel/invariance.lean b/src/probability/kernel/invariance.lean index 75ccfd8c63442..a33a340bb8e74 100644 --- a/src/probability/kernel/invariance.lean +++ b/src/probability/kernel/invariance.lean @@ -8,20 +8,17 @@ import probability.kernel.composition /-! # Invariance of measures along a kernel -We define the push-forward of a measure along a kernel which results in another measure. In the -case that the push-forward measure is the same as the original measure, we say that the measure is -invariant with respect to the kernel. +We say that a measure `μ` is invariant with respect to a kernel `κ` if its push-forward along the +kernel `μ.bind κ` is the same measure. ## Main definitions -* `probability_theory.kernel.map_measure`: the push-forward of a measure along a kernel. * `probability_theory.kernel.invariant`: invariance of a given measure with respect to a kernel. ## Useful lemmas -* `probability_theory.kernel.comp_apply_eq_map_measure`, - `probability_theory.kernel.const_map_measure_eq_comp_const`, and - `probability_theory.kernel.comp_const_apply_eq_map_measure` established the relationship between +* `probability_theory.kernel.const_bind_eq_comp_const`, and + `probability_theory.kernel.comp_const_apply_eq_bind` established the relationship between the push-forward measure and the composition of kernels. -/ @@ -41,87 +38,35 @@ namespace kernel /-! ### Push-forward of measures along a kernel -/ -/-- The push-forward of a measure along a kernel. -/ -noncomputable -def map_measure (κ : kernel α β) (μ : measure α) : - measure β := -measure.of_measurable (λ s hs, ∫⁻ x, κ x s ∂μ) - (by simp only [measure_empty, measure_theory.lintegral_const, zero_mul]) - begin - intros f hf₁ hf₂, - simp_rw [measure_Union hf₂ hf₁, - lintegral_tsum (λ i, (kernel.measurable_coe κ (hf₁ i)).ae_measurable)], - end - -@[simp] -lemma map_measure_apply (κ : kernel α β) (μ : measure α) {s : set β} (hs : measurable_set s) : - map_measure κ μ s = ∫⁻ x, κ x s ∂μ := -by rw [map_measure, measure.of_measurable_apply s hs] - @[simp] -lemma map_measure_zero (κ : kernel α β) : map_measure κ 0 = 0 := +lemma bind_add (μ ν : measure α) (κ : kernel α β) : + (μ + ν).bind κ = μ.bind κ + ν.bind κ := begin ext1 s hs, - rw [map_measure_apply κ 0 hs, lintegral_zero_measure, measure.coe_zero, pi.zero_apply], + rw [measure.bind_apply hs (kernel.measurable _), lintegral_add_measure, measure.coe_add, + pi.add_apply, measure.bind_apply hs (kernel.measurable _), + measure.bind_apply hs (kernel.measurable _)], end @[simp] -lemma map_measure_add (κ : kernel α β) (μ ν : measure α) : - map_measure κ (μ + ν) = map_measure κ μ + map_measure κ ν := +lemma bind_smul (κ : kernel α β) (μ : measure α) (r : ℝ≥0∞) : + (r • μ).bind κ = r • μ.bind κ := begin ext1 s hs, - rw [map_measure_apply κ (μ + ν) hs, lintegral_add_measure, measure.coe_add, pi.add_apply, - map_measure_apply κ μ hs, map_measure_apply κ ν hs], + rw [measure.bind_apply hs (kernel.measurable _), lintegral_smul_measure, measure.coe_smul, + pi.smul_apply, measure.bind_apply hs (kernel.measurable _), smul_eq_mul], end -@[simp] -lemma map_measure_smul (κ : kernel α β) (μ : measure α) (r : ℝ≥0∞) : - map_measure κ (r • μ) = r • map_measure κ μ := +lemma const_bind_eq_comp_const (κ : kernel α β) (μ : measure α) : + const α (μ.bind κ) = κ ∘ₖ const α μ := begin - ext1 s hs, - rw [map_measure_apply κ (r • μ) hs, lintegral_smul_measure, measure.coe_smul, pi.smul_apply, - map_measure_apply κ μ hs, smul_eq_mul], + ext a s hs : 2, + simp_rw [comp_apply' _ _ _ hs, const_apply, measure.bind_apply hs (kernel.measurable _)], end -include mγ - -lemma comp_apply_eq_map_measure - (η : kernel β γ) [is_s_finite_kernel η] (κ : kernel α β) [is_s_finite_kernel κ] (a : α) : - (η ∘ₖ κ) a = map_measure η (κ a) := -begin - ext1 s hs, - rw [comp_apply η κ a hs, map_measure_apply η _ hs], -end - -omit mγ - -lemma const_map_measure_eq_comp_const (κ : kernel α β) [is_s_finite_kernel κ] - (μ : measure α) [is_finite_measure μ] : - const α (map_measure κ μ) = κ ∘ₖ const α μ := -begin - ext1 a, ext1 s hs, - rw [const_apply, map_measure_apply _ _ hs, comp_apply _ _ _ hs, const_apply], -end - -lemma comp_const_apply_eq_map_measure (κ : kernel α β) [is_s_finite_kernel κ] - (μ : measure α) [is_finite_measure μ] (a : α) : - (κ ∘ₖ const α μ) a = map_measure κ μ := -by rw [← const_apply (map_measure κ μ) a, const_map_measure_eq_comp_const κ μ] - -lemma lintegral_map_measure - (κ : kernel α β) [is_s_finite_kernel κ] (μ : measure α) [is_finite_measure μ] - {f : β → ℝ≥0∞} (hf : measurable f) : - ∫⁻ b, f b ∂(map_measure κ μ) = ∫⁻ a, ∫⁻ b, f b ∂(κ a) ∂μ := -begin - by_cases hα : nonempty α, - { have := const_apply μ hα.some, - swap, apply_instance, - conv_rhs { rw [← this] }, - rw [← lintegral_comp _ _ _ hf, ← comp_const_apply_eq_map_measure κ μ hα.some] }, - { haveI := not_nonempty_iff.1 hα, - rw [μ.eq_zero_of_is_empty, map_measure_zero, lintegral_zero_measure, - lintegral_zero_measure] } -end +lemma comp_const_apply_eq_bind (κ : kernel α β) (μ : measure α) (a : α) : + (κ ∘ₖ const α μ) a = μ.bind κ := +by rw [← const_apply (μ.bind κ) a, const_bind_eq_comp_const κ μ] omit mβ @@ -130,24 +75,22 @@ omit mβ /-- A measure `μ` is invariant with respect to the kernel `κ` if the push-forward measure of `μ` along `κ` equals `μ`. -/ def invariant (κ : kernel α α) (μ : measure α) : Prop := -map_measure κ μ = μ +μ.bind κ = μ variables {κ η : kernel α α} {μ : measure α} -lemma invariant.def (hκ : invariant κ μ) : map_measure κ μ = μ := hκ +lemma invariant.def (hκ : invariant κ μ) : μ.bind κ = μ := hκ -lemma invariant.comp_const [is_s_finite_kernel κ] [is_finite_measure μ] - (hκ : invariant κ μ) : (κ ∘ₖ const α μ) = const α μ := -by rw [← const_map_measure_eq_comp_const κ μ, hκ.def] +lemma invariant.comp_const (hκ : invariant κ μ) : κ ∘ₖ const α μ = const α μ := +by rw [← const_bind_eq_comp_const κ μ, hκ.def] -lemma invariant.comp [is_s_finite_kernel κ] [is_s_finite_kernel η] [is_finite_measure μ] - (hκ : invariant κ μ) (hη : invariant η μ) : invariant (κ ∘ₖ η) μ := +lemma invariant.comp [is_s_finite_kernel κ] (hκ : invariant κ μ) (hη : invariant η μ) : + invariant (κ ∘ₖ η) μ := begin - by_cases hα : nonempty α, - { simp_rw [invariant, ← comp_const_apply_eq_map_measure (κ ∘ₖ η) μ hα.some, comp_assoc, + casesI is_empty_or_nonempty α with _ hα, + { exact subsingleton.elim _ _ }, + { simp_rw [invariant, ← comp_const_apply_eq_bind (κ ∘ₖ η) μ hα.some, comp_assoc, hη.comp_const, hκ.comp_const, const_apply] }, - { haveI := not_nonempty_iff.1 hα, - exact subsingleton.elim _ _ }, end end kernel From e354e865255654389cc46e6032160238df2e0f40 Mon Sep 17 00:00:00 2001 From: Floris van Doorn Date: Wed, 7 Jun 2023 13:00:51 +0000 Subject: [PATCH 1155/1291] feat(geometry/manifold): lemmas from the sphere eversion project (#18877) * Also adds a new library note `comp_of_eq lemmas` about how (I think) we should better formulate composition lemmas of properties of functions. * About the library note `comp_of_eq lemmas`: exactly the same problems happen in Lean 4. * renamings ``` smooth_iff_proj_smooth -> smooth_prod_iff differentiable_at.fderiv_within_prod -> differentiable_within_at.fderiv_within_prod ``` * We add a `path_connected_space` instance of the tangent space. This instance is sufficient to compile sphere-eversion, without any `normed_space` instances on the tangent space (which are not the canonical structure on the tangent space). * From the sphere eversion project --- src/analysis/calculus/fderiv/prod.lean | 2 +- src/geometry/manifold/algebra/monoid.lean | 11 + src/geometry/manifold/cont_mdiff.lean | 70 ++++- src/geometry/manifold/cont_mdiff_map.lean | 14 + src/geometry/manifold/cont_mdiff_mfderiv.lean | 21 ++ src/geometry/manifold/diffeomorph.lean | 5 + src/geometry/manifold/mfderiv.lean | 244 +++++++++++++++++- .../manifold/vector_bundle/tangent.lean | 15 ++ src/topology/algebra/module/basic.lean | 13 + src/topology/basic.lean | 34 +++ test/continuity.lean | 11 + 11 files changed, 424 insertions(+), 16 deletions(-) diff --git a/src/analysis/calculus/fderiv/prod.lean b/src/analysis/calculus/fderiv/prod.lean index 007fba3bc694d..d91f9ef1f8181 100644 --- a/src/analysis/calculus/fderiv/prod.lean +++ b/src/analysis/calculus/fderiv/prod.lean @@ -97,7 +97,7 @@ lemma differentiable_at.fderiv_prod fderiv 𝕜 (λx:E, (f₁ x, f₂ x)) x = (fderiv 𝕜 f₁ x).prod (fderiv 𝕜 f₂ x) := (hf₁.has_fderiv_at.prod hf₂.has_fderiv_at).fderiv -lemma differentiable_at.fderiv_within_prod +lemma differentiable_within_at.fderiv_within_prod (hf₁ : differentiable_within_at 𝕜 f₁ s x) (hf₂ : differentiable_within_at 𝕜 f₂ s x) (hxs : unique_diff_within_at 𝕜 s x) : fderiv_within 𝕜 (λx:E, (f₁ x, f₂ x)) s x = diff --git a/src/geometry/manifold/algebra/monoid.lean b/src/geometry/manifold/algebra/monoid.lean index 64c5c72065b9b..79d38983c7054 100644 --- a/src/geometry/manifold/algebra/monoid.lean +++ b/src/geometry/manifold/algebra/monoid.lean @@ -370,3 +370,14 @@ lemma smooth_finprod_cond (hc : ∀ i, p i → smooth I' I (f i)) cont_mdiff_finprod_cond hc hf end comm_monoid + +section + +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] +{E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] + +instance has_smooth_add_self : has_smooth_add 𝓘(𝕜, E) E := +⟨by { convert cont_diff_add.cont_mdiff, exact model_with_corners_self_prod.symm, + exact charted_space_self_prod }⟩ + +end diff --git a/src/geometry/manifold/cont_mdiff.lean b/src/geometry/manifold/cont_mdiff.lean index c0e101cfb1241..72c43255f6672 100644 --- a/src/geometry/manifold/cont_mdiff.lean +++ b/src/geometry/manifold/cont_mdiff.lean @@ -925,6 +925,13 @@ begin { simp only [written_in_ext_chart_at, (∘), mem_ext_chart_source, e.left_inv, e'.left_inv] } end +/-- See note [comp_of_eq lemmas] -/ +lemma cont_mdiff_within_at.comp_of_eq {t : set M'} {g : M' → M''} {x : M} {y : M'} + (hg : cont_mdiff_within_at I' I'' n g t y) (hf : cont_mdiff_within_at I I' n f s x) + (st : maps_to f s t) (hx : f x = y) : + cont_mdiff_within_at I I'' n (g ∘ f) s x := +by { subst hx, exact hg.comp x hf st } + /-- The composition of `C^∞` functions within domains at points is `C^∞`. -/ lemma smooth_within_at.comp {t : set M'} {g : M' → M''} (x : M) (hg : smooth_within_at I' I'' g t (f x)) @@ -1004,6 +1011,13 @@ lemma cont_mdiff_at.comp {g : M' → M''} (x : M) cont_mdiff_at I I'' n (g ∘ f) x := hg.comp x hf (maps_to_univ _ _) +/-- See note [comp_of_eq lemmas] -/ +lemma cont_mdiff_at.comp_of_eq {g : M' → M''} {x : M} {y : M'} + (hg : cont_mdiff_at I' I'' n g y) (hf : cont_mdiff_at I I' n f x) (hx : f x = y) : + cont_mdiff_at I I'' n (g ∘ f) x := +by { subst hx, exact hg.comp x hf } + + /-- The composition of `C^∞` functions at points is `C^∞`. -/ lemma smooth_at.comp {g : M' → M''} (x : M) (hg : smooth_at I' I'' g (f x)) (hf : smooth_at I I' f x) : @@ -1045,6 +1059,14 @@ begin simp_rw [function.comp_apply, I.left_inv, id_def] end +lemma cont_mdiff_on_model_symm : cont_mdiff_on 𝓘(𝕜, E) I n I.symm (range I) := +begin + rw [cont_mdiff_on_iff], + refine ⟨I.continuous_on_symm, λ x y, _⟩, + simp only with mfld_simps, + exact cont_diff_on_id.congr (λ x', I.right_inv) +end + include Is /-- An atlas member is `C^n` for any `n`. -/ @@ -1092,6 +1114,22 @@ lemma cont_mdiff_on_ext_chart_at : cont_mdiff_on I 𝓘(𝕜, E) n (ext_chart_at I x) (chart_at H x).source := λ x' hx', (cont_mdiff_at_ext_chart_at' hx').cont_mdiff_within_at +lemma cont_mdiff_on_extend_symm (he : e ∈ maximal_atlas I M) : + cont_mdiff_on 𝓘(𝕜, E) I n (e.extend I).symm (I '' e.target) := +begin + have h2 := cont_mdiff_on_symm_of_mem_maximal_atlas he, + refine h2.comp (cont_mdiff_on_model_symm.mono $ image_subset_range _ _) _, + simp_rw [image_subset_iff, local_equiv.restr_coe_symm, I.to_local_equiv_coe_symm, + preimage_preimage, I.left_inv, preimage_id'] +end + +lemma cont_mdiff_on_ext_chart_at_symm (x : M) : + cont_mdiff_on 𝓘(𝕜, E) I n (ext_chart_at I x).symm (ext_chart_at I x).target := +begin + convert cont_mdiff_on_extend_symm (chart_mem_maximal_atlas I x), + rw [ext_chart_at_target, I.image_eq] +end + omit Is /-- An element of `cont_diff_groupoid ⊤ I` is `C^n` for any `n`. -/ @@ -1490,18 +1528,6 @@ lemma smooth.snd {f : N → M × M'} (hf : smooth J (I.prod I') f) : smooth J I' (λ x, (f x).2) := smooth_snd.comp hf -lemma smooth_iff_proj_smooth {f : M → M' × N'} : - (smooth I (I'.prod J') f) ↔ (smooth I I' (prod.fst ∘ f)) ∧ (smooth I J' (prod.snd ∘ f)) := -begin - split, - { intro h, exact ⟨smooth_fst.comp h, smooth_snd.comp h⟩ }, - { rintro ⟨h_fst, h_snd⟩, simpa only [prod.mk.eta] using h_fst.prod_mk h_snd, } -end - -lemma smooth_prod_assoc : - smooth ((I.prod I').prod J) (I.prod (I'.prod J)) (λ x : (M × M') × N, (x.1.1, x.1.2, x.2)) := -smooth_fst.fst.prod_mk $ smooth_fst.snd.prod_mk smooth_snd - end projections lemma cont_mdiff_within_at_prod_iff (f : M → M' × N') {s : set M} {x : M} : @@ -1513,7 +1539,25 @@ by { refine ⟨λ h, ⟨h.fst, h.snd⟩, λ h, _⟩, simpa only [prod.mk.eta] us lemma cont_mdiff_at_prod_iff (f : M → M' × N') {x : M} : cont_mdiff_at I (I'.prod J') n f x ↔ cont_mdiff_at I I' n (prod.fst ∘ f) x ∧ cont_mdiff_at I J' n (prod.snd ∘ f) x := -by { simp_rw [← cont_mdiff_within_at_univ], exact cont_mdiff_within_at_prod_iff f } +by simp_rw [← cont_mdiff_within_at_univ, cont_mdiff_within_at_prod_iff] + +lemma cont_mdiff_prod_iff (f : M → M' × N') : + cont_mdiff I (I'.prod J') n f ↔ + cont_mdiff I I' n (prod.fst ∘ f) ∧ cont_mdiff I J' n (prod.snd ∘ f) := +⟨λ h, ⟨h.fst, h.snd⟩, λ h, by { convert h.1.prod_mk h.2, ext; refl }⟩ + +lemma smooth_at_prod_iff (f : M → M' × N') {x : M} : + smooth_at I (I'.prod J') f x ↔ + smooth_at I I' (prod.fst ∘ f) x ∧ smooth_at I J' (prod.snd ∘ f) x := +cont_mdiff_at_prod_iff f + +lemma smooth_prod_iff (f : M → M' × N') : + smooth I (I'.prod J') f ↔ smooth I I' (prod.fst ∘ f) ∧ smooth I J' (prod.snd ∘ f) := +cont_mdiff_prod_iff f + +lemma smooth_prod_assoc : + smooth ((I.prod I').prod J) (I.prod (I'.prod J)) (λ x : (M × M') × N, (x.1.1, x.1.2, x.2)) := +smooth_fst.fst.prod_mk $ smooth_fst.snd.prod_mk smooth_snd section prod_map diff --git a/src/geometry/manifold/cont_mdiff_map.lean b/src/geometry/manifold/cont_mdiff_map.lean index b064885bcbda5..1fce5586befc0 100644 --- a/src/geometry/manifold/cont_mdiff_map.lean +++ b/src/geometry/manifold/cont_mdiff_map.lean @@ -26,6 +26,10 @@ variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] {H'' : Type*} [topological_space H''] {I'' : model_with_corners 𝕜 E'' H''} {M'' : Type*} [topological_space M''] [charted_space H'' M''] +-- declare a manifold `N` over the pair `(F, G)`. +{F : Type*} [normed_add_comm_group F] [normed_space 𝕜 F] +{G : Type*} [topological_space G] {J : model_with_corners 𝕜 F G} +{N : Type*} [topological_space N] [charted_space G N] (n : ℕ∞) /-- Bundled `n` times continuously differentiable maps. -/ @@ -94,6 +98,16 @@ instance [inhabited M'] : inhabited C^n⟮I, M; I', M'⟯ := /-- Constant map as a smooth map -/ def const (y : M') : C^n⟮I, M; I', M'⟯ := ⟨λ x, y, cont_mdiff_const⟩ +/-- The first projection of a product, as a smooth map. -/ +def fst : C^n⟮I.prod I', M × M'; I, M⟯ := ⟨prod.fst, cont_mdiff_fst⟩ + +/-- The second projection of a product, as a smooth map. -/ +def snd : C^n⟮I.prod I', M × M'; I', M'⟯ := ⟨prod.snd, cont_mdiff_snd⟩ + +/-- Given two smooth maps `f` and `g`, this is the smooth map `x ↦ (f x, g x)`. -/ +def prod_mk (f : C^n⟮J, N; I, M⟯) (g : C^n⟮J, N; I', M'⟯) : C^n⟮J, N; I.prod I', M × M'⟯ := +⟨λ x, (f x, g x), f.2.prod_mk g.2⟩ + end cont_mdiff_map instance continuous_linear_map.has_coe_to_cont_mdiff_map : diff --git a/src/geometry/manifold/cont_mdiff_mfderiv.lean b/src/geometry/manifold/cont_mdiff_mfderiv.lean index 3ba92b1b57c6d..048675244325c 100644 --- a/src/geometry/manifold/cont_mdiff_mfderiv.lean +++ b/src/geometry/manifold/cont_mdiff_mfderiv.lean @@ -214,6 +214,27 @@ begin exact this.mfderiv (λ x, f) id cont_mdiff_at_id hmn, end +include Js +/-- The function that sends `x` to the `y`-derivative of `f(x,y)` at `g(x)` applied to `g₂(x)` is +`C^n` at `x₀`, where the derivative is taken as a continuous linear map. +We have to assume that `f` is `C^(n+1)` at `(x₀, g(x₀))` and `g` is `C^n` at `x₀`. +We have to insert a coordinate change from `x₀` to `g₁(x)` to make the derivative sensible. + +This is similar to `cont_mdiff_at.mfderiv`, but where the continuous linear map is applied to a +(variable) vector. +-/ +lemma cont_mdiff_at.mfderiv_apply {x₀ : N'} (f : N → M → M') (g : N → M) (g₁ : N' → N) + (g₂ : N' → E) + (hf : cont_mdiff_at (J.prod I) I' n (function.uncurry f) (g₁ x₀, g (g₁ x₀))) + (hg : cont_mdiff_at J I m g (g₁ x₀)) + (hg₁ : cont_mdiff_at J' J m g₁ x₀) + (hg₂ : cont_mdiff_at J' 𝓘(𝕜, E) m g₂ x₀) (hmn : m + 1 ≤ n) : + cont_mdiff_at J' 𝓘(𝕜, E') m + (λ x, in_tangent_coordinates I I' g (λ x, f x (g x)) (λ x, mfderiv I I' (f x) (g x)) + (g₁ x₀) (g₁ x) (g₂ x)) + x₀ := +((hf.mfderiv f g hg hmn).comp_of_eq hg₁ rfl).clm_apply hg₂ + end mfderiv /-! ### The tangent map of a smooth function is smooth -/ diff --git a/src/geometry/manifold/diffeomorph.lean b/src/geometry/manifold/diffeomorph.lean index 49a538a391b4a..4a1a820e6dfda 100644 --- a/src/geometry/manifold/diffeomorph.lean +++ b/src/geometry/manifold/diffeomorph.lean @@ -126,6 +126,11 @@ equiv.coe_fn_injective.comp to_equiv_injective @[ext] lemma ext {h h' : M ≃ₘ^n⟮I, I'⟯ M'} (Heq : ∀ x, h x = h' x) : h = h' := coe_fn_injective $ funext Heq +instance : continuous_map_class (M ≃ₘ⟮I, J⟯ N) M N := +{ coe := coe_fn, + coe_injective' := coe_fn_injective, + map_continuous := λ f, f.continuous } + section variables (M I n) diff --git a/src/geometry/manifold/mfderiv.lean b/src/geometry/manifold/mfderiv.lean index 971283f6f685a..67669519a2b97 100644 --- a/src/geometry/manifold/mfderiv.lean +++ b/src/geometry/manifold/mfderiv.lean @@ -94,7 +94,7 @@ Derivative, manifold noncomputable theory open_locale classical topology manifold bundle -open set +open set bundle universe u @@ -696,6 +696,54 @@ end omit Is I's +lemma mdifferentiable_within_at.prod_mk {f : M → M'} {g : M → M''} + (hf : mdifferentiable_within_at I I' f s x) + (hg : mdifferentiable_within_at I I'' g s x) : + mdifferentiable_within_at I (I'.prod I'') (λ x, (f x, g x)) s x := +⟨hf.1.prod hg.1, hf.2.prod hg.2⟩ + +lemma mdifferentiable_at.prod_mk {f : M → M'} {g : M → M''} + (hf : mdifferentiable_at I I' f x) + (hg : mdifferentiable_at I I'' g x) : + mdifferentiable_at I (I'.prod I'') (λ x, (f x, g x)) x := +⟨hf.1.prod hg.1, hf.2.prod hg.2⟩ + +lemma mdifferentiable_on.prod_mk {f : M → M'} {g : M → M''} + (hf : mdifferentiable_on I I' f s) + (hg : mdifferentiable_on I I'' g s) : + mdifferentiable_on I (I'.prod I'') (λ x, (f x, g x)) s := +λ x hx, (hf x hx).prod_mk (hg x hx) + +lemma mdifferentiable.prod_mk {f : M → M'} {g : M → M''} + (hf : mdifferentiable I I' f) + (hg : mdifferentiable I I'' g) : + mdifferentiable I (I'.prod I'') (λ x, (f x, g x)) := +λ x, (hf x).prod_mk (hg x) + +lemma mdifferentiable_within_at.prod_mk_space {f : M → E'} {g : M → E''} + (hf : mdifferentiable_within_at I 𝓘(𝕜, E') f s x) + (hg : mdifferentiable_within_at I 𝓘(𝕜, E'') g s x) : + mdifferentiable_within_at I 𝓘(𝕜, E' × E'') (λ x, (f x, g x)) s x := +⟨hf.1.prod hg.1, hf.2.prod hg.2⟩ + +lemma mdifferentiable_at.prod_mk_space {f : M → E'} {g : M → E''} + (hf : mdifferentiable_at I 𝓘(𝕜, E') f x) + (hg : mdifferentiable_at I 𝓘(𝕜, E'') g x) : + mdifferentiable_at I 𝓘(𝕜, E' × E'') (λ x, (f x, g x)) x := +⟨hf.1.prod hg.1, hf.2.prod hg.2⟩ + +lemma mdifferentiable_on.prod_mk_space {f : M → E'} {g : M → E''} + (hf : mdifferentiable_on I 𝓘(𝕜, E') f s) + (hg : mdifferentiable_on I 𝓘(𝕜, E'') g s) : + mdifferentiable_on I 𝓘(𝕜, E' × E'') (λ x, (f x, g x)) s := +λ x hx, (hf x hx).prod_mk_space (hg x hx) + +lemma mdifferentiable.prod_mk_space {f : M → E'} {g : M → E''} + (hf : mdifferentiable I 𝓘(𝕜, E') f) + (hg : mdifferentiable I 𝓘(𝕜, E'') g) : + mdifferentiable I 𝓘(𝕜, E' × E'') (λ x, (f x, g x)) := +λ x, (hf x).prod_mk_space (hg x) + /-! ### Congruence lemmas for derivatives on manifolds -/ lemma has_mfderiv_within_at.congr_of_eventually_eq (h : has_mfderiv_within_at I I' f s x f') @@ -805,6 +853,18 @@ begin exact hL.mfderiv_within_eq (unique_mdiff_within_at_univ I) A end +/-- A congruence lemma for `mfderiv`, (ab)using the fact that `tangent_space I' (f x)` is +definitionally equal to `E'`. -/ +lemma mfderiv_congr_point {x' : M} (h : x = x') : + @eq (E →L[𝕜] E') (mfderiv I I' f x) (mfderiv I I' f x') := +by subst h + +/-- A congruence lemma for `mfderiv`, (ab)using the fact that `tangent_space I' (f x)` is +definitionally equal to `E'`. -/ +lemma mfderiv_congr {f' : M → M'} (h : f = f') : + @eq (E →L[𝕜] E') (mfderiv I I' f x) (mfderiv I I' f' x) := +by subst h + /-! ### Composition lemmas -/ omit Is I's @@ -905,6 +965,11 @@ begin exact has_mfderiv_at.comp x hg.has_mfderiv_at hf.has_mfderiv_at end +lemma mfderiv_comp_of_eq {x : M} {y : M'} + (hg : mdifferentiable_at I' I'' g y) (hf : mdifferentiable_at I I' f x) (hy : f x = y) : + mfderiv I I'' (g ∘ f) x = (mfderiv I' I'' g (f x)).comp (mfderiv I I' f x) := +by { subst hy, exact mfderiv_comp x hg hf } + lemma mdifferentiable_on.comp (hg : mdifferentiable_on I' I'' g u) (hf : mdifferentiable_on I I' f s) (st : s ⊆ f ⁻¹' u) : mdifferentiable_on I I'' (g ∘ f) s := @@ -1063,6 +1128,9 @@ variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] {E' : Type*} [normed_add_comm_group E'] [normed_space 𝕜 E'] {H' : Type*} [topological_space H'] (I' : model_with_corners 𝕜 E' H') {M' : Type*} [topological_space M'] [charted_space H' M'] [smooth_manifold_with_corners I' M'] +{E'' : Type*} [normed_add_comm_group E''] [normed_space 𝕜 E''] +{H'' : Type*} [topological_space H''] (I'' : model_with_corners 𝕜 E'' H'') +{M'' : Type*} [topological_space M''] [charted_space H'' M''] [smooth_manifold_with_corners I'' M''] namespace continuous_linear_map @@ -1135,7 +1203,7 @@ section id lemma has_mfderiv_at_id (x : M) : has_mfderiv_at I I (@_root_.id M) x (continuous_linear_map.id 𝕜 (tangent_space I x)) := begin - refine ⟨continuous_id.continuous_at, _⟩, + refine ⟨continuous_at_id, _⟩, have : ∀ᶠ y in 𝓝[range I] ((ext_chart_at I x) x), ((ext_chart_at I x) ∘ (ext_chart_at I x).symm) y = id y, { apply filter.mem_of_superset (ext_chart_at_target_mem_nhds_within I x), @@ -1227,6 +1295,178 @@ lemma mfderiv_within_const (hxs : unique_mdiff_within_at I s x) : end const +section prod +/-! Operations on the product of two manifolds-/ + +lemma has_mfderiv_at_fst (x : M × M') : + has_mfderiv_at (I.prod I') I prod.fst x + (continuous_linear_map.fst 𝕜 (tangent_space I x.1) (tangent_space I' x.2)) := +begin + refine ⟨continuous_fst.continuous_at, _⟩, + have : ∀ᶠ y in 𝓝[range (I.prod I')] (ext_chart_at (I.prod I') x x), + ((ext_chart_at I x.1) ∘ prod.fst ∘ (ext_chart_at (I.prod I') x).symm) y = y.1, + { apply filter.mem_of_superset (ext_chart_at_target_mem_nhds_within (I.prod I') x), + mfld_set_tac }, + apply has_fderiv_within_at.congr_of_eventually_eq has_fderiv_within_at_fst this, + simp only with mfld_simps +end + +theorem has_mfderiv_within_at_fst (s : set (M × M')) (x : M × M') : + has_mfderiv_within_at (I.prod I') I prod.fst s x + (continuous_linear_map.fst 𝕜 (tangent_space I x.1) (tangent_space I' x.2)) := +(has_mfderiv_at_fst I I' x).has_mfderiv_within_at + +lemma mdifferentiable_at_fst {x : M × M'} : mdifferentiable_at (I.prod I') I prod.fst x := +(has_mfderiv_at_fst I I' x).mdifferentiable_at + +lemma mdifferentiable_within_at_fst {s : set (M × M')} {x : M × M'} : + mdifferentiable_within_at (I.prod I') I prod.fst s x := +(mdifferentiable_at_fst I I').mdifferentiable_within_at + +lemma mdifferentiable_fst : mdifferentiable (I.prod I') I (prod.fst : M × M' → M) := +λx, mdifferentiable_at_fst I I' + +lemma mdifferentiable_on_fst {s : set (M × M')} : mdifferentiable_on (I.prod I') I prod.fst s := +(mdifferentiable_fst I I').mdifferentiable_on + +@[simp, mfld_simps] lemma mfderiv_fst {x : M × M'} : + mfderiv (I.prod I') I prod.fst x = + continuous_linear_map.fst 𝕜 (tangent_space I x.1) (tangent_space I' x.2) := +(has_mfderiv_at_fst I I' x).mfderiv + +lemma mfderiv_within_fst {s : set (M × M')} {x : M × M'} + (hxs : unique_mdiff_within_at (I.prod I') s x) : + mfderiv_within (I.prod I') I prod.fst s x = + continuous_linear_map.fst 𝕜 (tangent_space I x.1) (tangent_space I' x.2) := +by { rw mdifferentiable.mfderiv_within (mdifferentiable_at_fst I I') hxs, exact mfderiv_fst I I' } + +@[simp, mfld_simps] lemma tangent_map_prod_fst {p : tangent_bundle (I.prod I') (M × M')} : + tangent_map (I.prod I') I prod.fst p = total_space_mk p.proj.1 p.2.1 := +by simp [tangent_map] + +lemma tangent_map_within_prod_fst {s : set (M × M')} {p : tangent_bundle (I.prod I') (M × M')} + (hs : unique_mdiff_within_at (I.prod I') s p.proj) : + tangent_map_within (I.prod I') I prod.fst s p = total_space_mk p.proj.1 p.2.1 := +begin + simp only [tangent_map_within], + rw mfderiv_within_fst, + { rcases p, refl }, + { exact hs } +end + +lemma has_mfderiv_at_snd (x : M × M') : + has_mfderiv_at (I.prod I') I' prod.snd x + (continuous_linear_map.snd 𝕜 (tangent_space I x.1) (tangent_space I' x.2)) := +begin + refine ⟨continuous_snd.continuous_at, _⟩, + have : ∀ᶠ y in 𝓝[range (I.prod I')] (ext_chart_at (I.prod I') x x), + ((ext_chart_at I' x.2) ∘ prod.snd ∘ (ext_chart_at (I.prod I') x).symm) y = y.2, + { apply filter.mem_of_superset (ext_chart_at_target_mem_nhds_within (I.prod I') x), + mfld_set_tac }, + apply has_fderiv_within_at.congr_of_eventually_eq has_fderiv_within_at_snd this, + simp only with mfld_simps +end + +theorem has_mfderiv_within_at_snd (s : set (M × M')) (x : M × M') : + has_mfderiv_within_at (I.prod I') I' prod.snd s x + (continuous_linear_map.snd 𝕜 (tangent_space I x.1) (tangent_space I' x.2)) := +(has_mfderiv_at_snd I I' x).has_mfderiv_within_at + +lemma mdifferentiable_at_snd {x : M × M'} : mdifferentiable_at (I.prod I') I' prod.snd x := +(has_mfderiv_at_snd I I' x).mdifferentiable_at + +lemma mdifferentiable_within_at_snd {s : set (M × M')} {x : M × M'} : + mdifferentiable_within_at (I.prod I') I' prod.snd s x := +(mdifferentiable_at_snd I I').mdifferentiable_within_at + +lemma mdifferentiable_snd : mdifferentiable (I.prod I') I' (prod.snd : M × M' → M') := +λx, mdifferentiable_at_snd I I' + +lemma mdifferentiable_on_snd {s : set (M × M')} : mdifferentiable_on (I.prod I') I' prod.snd s := +(mdifferentiable_snd I I').mdifferentiable_on + +@[simp, mfld_simps] lemma mfderiv_snd {x : M × M'} : + mfderiv (I.prod I') I' prod.snd x = + continuous_linear_map.snd 𝕜 (tangent_space I x.1) (tangent_space I' x.2) := +(has_mfderiv_at_snd I I' x).mfderiv + +lemma mfderiv_within_snd {s : set (M × M')} {x : M × M'} + (hxs : unique_mdiff_within_at (I.prod I') s x) : + mfderiv_within (I.prod I') I' prod.snd s x = + continuous_linear_map.snd 𝕜 (tangent_space I x.1) (tangent_space I' x.2) := +by { rw mdifferentiable.mfderiv_within (mdifferentiable_at_snd I I') hxs, exact mfderiv_snd I I' } + +@[simp, mfld_simps] lemma tangent_map_prod_snd {p : tangent_bundle (I.prod I') (M × M')} : + tangent_map (I.prod I') I' prod.snd p = total_space_mk p.proj.2 p.2.2 := +by simp [tangent_map] + +lemma tangent_map_within_prod_snd {s : set (M × M')} {p : tangent_bundle (I.prod I') (M × M')} + (hs : unique_mdiff_within_at (I.prod I') s p.proj) : + tangent_map_within (I.prod I') I' prod.snd s p = total_space_mk p.proj.2 p.2.2 := +begin + simp only [tangent_map_within], + rw mfderiv_within_snd, + { rcases p, refl }, + { exact hs } +end + +variables {I I' I''} +lemma mdifferentiable_at.mfderiv_prod {f : M → M'} {g : M → M''} {x : M} + (hf : mdifferentiable_at I I' f x) + (hg : mdifferentiable_at I I'' g x) : + mfderiv I (I'.prod I'') (λ x, (f x, g x)) x = (mfderiv I I' f x).prod (mfderiv I I'' g x) := +begin + classical, + simp_rw [mfderiv, if_pos (hf.prod_mk hg), if_pos hf, if_pos hg], + exact hf.2.fderiv_within_prod hg.2 (I.unique_diff _ (mem_range_self _)) +end + +variables (I I' I'') + +lemma mfderiv_prod_left {x₀ : M} {y₀ : M'} : + mfderiv I (I.prod I') (λ x, (x, y₀)) x₀ = + continuous_linear_map.inl 𝕜 (tangent_space I x₀) (tangent_space I' y₀) := +begin + refine ((mdifferentiable_at_id I).mfderiv_prod (mdifferentiable_at_const I I')).trans _, + rw [mfderiv_id, mfderiv_const, continuous_linear_map.inl] +end + +lemma mfderiv_prod_right {x₀ : M} {y₀ : M'} : + mfderiv I' (I.prod I') (λ y, (x₀, y)) y₀ = + continuous_linear_map.inr 𝕜 (tangent_space I x₀) (tangent_space I' y₀) := +begin + refine ((mdifferentiable_at_const I' I).mfderiv_prod (mdifferentiable_at_id I')).trans _, + rw [mfderiv_id, mfderiv_const, continuous_linear_map.inr] +end + +/-- The total derivative of a function in two variables is the sum of the partial derivatives. + Note that to state this (without casts) we need to be able to see through the definition of + `tangent_space`. -/ +lemma mfderiv_prod_eq_add {f : M × M' → M''} {p : M × M'} + (hf : mdifferentiable_at (I.prod I') I'' f p) : + mfderiv (I.prod I') I'' f p = + (show E × E' →L[𝕜] E'', from mfderiv (I.prod I') I'' (λ (z : M × M'), f (z.1, p.2)) p + + mfderiv (I.prod I') I'' (λ (z : M × M'), f (p.1, z.2)) p) := +begin + dsimp only, + rw [← @prod.mk.eta _ _ p] at hf, + rw [mfderiv_comp_of_eq hf ((mdifferentiable_at_fst I I').prod_mk (mdifferentiable_at_const _ _)) + rfl, + mfderiv_comp_of_eq hf ((mdifferentiable_at_const _ _).prod_mk (mdifferentiable_at_snd I I')) + rfl, + ← continuous_linear_map.comp_add, + (mdifferentiable_at_fst I I').mfderiv_prod (mdifferentiable_at_const (I.prod I') I'), + (mdifferentiable_at_const (I.prod I') I).mfderiv_prod (mdifferentiable_at_snd I I'), + mfderiv_fst, mfderiv_snd, mfderiv_const, mfderiv_const], + symmetry, + convert continuous_linear_map.comp_id _, + { exact continuous_linear_map.coprod_inl_inr }, + simp_rw [prod.mk.eta], + all_goals { apply_instance } +end + +end prod + section arithmetic /-! #### Arithmetic diff --git a/src/geometry/manifold/vector_bundle/tangent.lean b/src/geometry/manifold/vector_bundle/tangent.lean index f8942eb407f51..f5cd27d8429a8 100644 --- a/src/geometry/manifold/vector_bundle/tangent.lean +++ b/src/geometry/manifold/vector_bundle/tangent.lean @@ -34,6 +34,8 @@ open_locale manifold topology bundle noncomputable theory +section general + variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] {E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] {E' : Type*} [normed_add_comm_group E'] [normed_space 𝕜 E'] @@ -148,6 +150,7 @@ variables {M} (x : M) instance : module 𝕜 (tangent_space I x) := by delta_instance tangent_space instance : inhabited (tangent_space I x) := ⟨0⟩ +instance {x : M} : has_continuous_add (tangent_space I x) := by delta_instance tangent_space end @@ -387,3 +390,15 @@ lemma in_tangent_coordinates_eq (f : N → M) (g : N → M') (ϕ : N → E →L[ (tangent_bundle_core I M).in_coordinates_eq (tangent_bundle_core I' M') (ϕ x) hx hy end in_tangent_coordinates + +end general + +section real + +variables {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] +{H : Type*} [topological_space H] {I : model_with_corners ℝ E H} +{M : Type*} [topological_space M] [charted_space H M] [smooth_manifold_with_corners I M] + +instance {x : M} : path_connected_space (tangent_space I x) := by delta_instance tangent_space + +end real diff --git a/src/topology/algebra/module/basic.lean b/src/topology/algebra/module/basic.lean index aafd5fefb87f1..5379fd79c140b 100644 --- a/src/topology/algebra/module/basic.lean +++ b/src/topology/algebra/module/basic.lean @@ -955,6 +955,19 @@ lemma range_coprod [module R₁ M₂] [module R₁ M₃] [has_continuous_add M range (f₁.coprod f₂) = range f₁ ⊔ range f₂ := linear_map.range_coprod _ _ +lemma comp_fst_add_comp_snd [module R₁ M₂] [module R₁ M₃] [has_continuous_add M₃] + (f : M₁ →L[R₁] M₃) (g : M₂ →L[R₁] M₃) : + f.comp (continuous_linear_map.fst R₁ M₁ M₂) + + g.comp (continuous_linear_map.snd R₁ M₁ M₂) = + f.coprod g := +rfl + + +lemma coprod_inl_inr [has_continuous_add M₁] [has_continuous_add M'₁] : + (continuous_linear_map.inl R₁ M₁ M'₁).coprod (continuous_linear_map.inr R₁ M₁ M'₁) = + continuous_linear_map.id R₁ (M₁ × M'₁) := +by { apply coe_injective, apply linear_map.coprod_inl_inr } + section variables {R S : Type*} [semiring R] [semiring S] [module R M₁] [module R M₂] [module R S] diff --git a/src/topology/basic.lean b/src/topology/basic.lean index 2eb7568ea7d7c..591b548277ee2 100644 --- a/src/topology/basic.lean +++ b/src/topology/basic.lean @@ -1320,6 +1320,12 @@ lemma continuous_at.comp {g : β → γ} {f : α → β} {x : α} continuous_at (g ∘ f) x := hg.comp hf +/-- See note [comp_of_eq lemmas] -/ +lemma continuous_at.comp_of_eq {g : β → γ} {f : α → β} {x : α} {y : β} + (hg : continuous_at g y) (hf : continuous_at f x) (hy : f x = y) : + continuous_at (g ∘ f) x := +by { subst hy, exact hg.comp hf } + lemma continuous.tendsto {f : α → β} (hf : continuous f) (x) : tendsto f (𝓝 x) (𝓝 (f x)) := ((nhds_basis_opens x).tendsto_iff $ nhds_basis_opens $ f x).2 $ @@ -1607,3 +1613,31 @@ With `continuous_at` you can be even more precise about what to prove in case of see e.g. `continuous_at.comp_div_cases`. -/ library_note "continuity lemma statement" + +/-- +Lean's elaborator has trouble elaborating applications of lemmas that state that the composition of +two functions satisfy some property at a point, like `continuous_at.comp` / `cont_diff_at.comp` and +`cont_mdiff_within_at.comp`. The reason is that a lemma like this looks like +`continuous_at g (f x) → continuous_at f x → continuous_at (g ∘ f) x`. +Since Lean's elaborator elaborates the arguments from left-to-right, when you write `hg.comp hf`, +the elaborator will try to figure out *both* `f` and `g` from the type of `hg`. It tries to figure +out `f` just from the point where `g` is continuous. For example, if `hg : continuous_at g (a, x)` +then the elaborator will assign `f` to the function `prod.mk a`, since in that case `f x = (a, x)`. +This is undesirable in most cases where `f` is not a variable. There are some ways to work around +this, for example by giving `f` explicitly, or to force Lean to elaborate `hf` before elaborating +`hg`, but this is annoying. +Another better solution is to reformulate composition lemmas to have the following shape +`continuous_at g y → continuous_at f x → f x = y → continuous_at (g ∘ f) x`. +This is even useful if the proof of `f x = y` is `rfl`. +The reason that this works better is because the type of `hg` doesn't mention `f`. +Only after elaborating the two `continuous_at` arguments, Lean will try to unify `f x` with `y`, +which is often easy after having chosen the correct functions for `f` and `g`. +Here is an example that shows the difference: +``` +example {x₀ : α} (f : α → α → β) (hf : continuous_at (function.uncurry f) (x₀, x₀)) : + continuous_at (λ x => f x x) x₀ := +-- hf.comp x (continuous_at_id.prod continuous_at_id) -- type mismatch +-- hf.comp_of_eq (continuous_at_id.prod continuous_at_id) rfl -- works +``` +-/ +library_note "comp_of_eq lemmas" diff --git a/test/continuity.lean b/test/continuity.lean index 6801d76bd80ae..ff2a2633c2acf 100644 --- a/test/continuity.lean +++ b/test/continuity.lean @@ -39,3 +39,14 @@ by guard_proof_term { continuity } -- ⊢ continuous complex.exp -- ⊢ continuous coe -- ⊢ continuous (λ (x : ℝ), ↑x) + + +/-! Some tests of the `comp_of_eq` lemmas -/ + +example {α β : Type*} [topological_space α] [topological_space β] {x₀ : α} (f : α → α → β) + (hf : continuous_at (function.uncurry f) (x₀, x₀)) : + continuous_at (λ x, f x x) x₀ := +begin + success_if_fail { exact hf.comp x (continuous_at_id.prod continuous_at_id) }, + exact hf.comp_of_eq (continuous_at_id.prod continuous_at_id) rfl +end From 31c24aa72e7b3e5ed97a8412470e904f82b81004 Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Wed, 7 Jun 2023 16:34:13 +0000 Subject: [PATCH 1156/1291] feat(algebra/star/basic): refactor `star_ordered_ring` to include `add_submonoid.closure` (#18854) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Per [Zulip](https://leanprover.zulipchat.com/#narrow/stream/116395-maths/topic/Star.20ordered.20ring), this refactors `star_ordered_ring` so that the condition `star_ordered_ring.nonneg_iff` is changed from `∀ r : R, 0 ≤ r ↔ ∃ s, r = star s * s` to something morally equivalent to `∀ x : R, 0 ≤ x ↔ x ∈ add_submonoid.closure (set.range (λ s : R, star s * s))`. In fact, we actually change the structure field `nonneg_iff` to `le_iff`, which characterizes `· ≤ ·` instead of just `0 ≤ ·`. When `R` is a `non_unital_ring`, there is effectively no change (see how we recover `star_ordered_ring.nonneg_iff` and also `star_ordered_ring.of_nonneg_iff`), but it gives a more useful and sensible condition when `R` is only a `non_unital_semiring`. For instance, now `conjugate_le_conjugate` holds for `non_unital_semiring`. There are essentially two reasons for this change. 1. It would be nice if things like `ℚ` could be `star_ordered_ring`s. This is a minor reason, but it should be a nice convenience. This instance is added in this PR in a new file. 2. Much more importantly, we want to declare the positive elements in a `star_ordered_ring` as an `add_submonoid`, but to accomplish this with the previous definition requires much more stringent type class assumptions (e.g., C⋆-algebras) and sophisticated machinery (the continuous functional calculus) in order to show that the sum of positive elements is positive. This change essentially allows us to defer that proof obligation to the settings where it will matter that a positive element really does have the form `star s * s`. We remark that even for C⋆-algebras, the fact that the sum of positive elements (i.e., those of the form `star s * s`) is positive is a deep result which was first shown in 1952 by [Fukamiya](http://www.sci.kumamoto-u.ac.jp/~kjm/BKS/kjmpdf/KJSM/v1-4-fukamiya.pdf), and then again in 1953 by [Kelley and Vaught](https://www.ams.org/journals/tran/1953-074-01/S0002-9947-1953-0054175-2/S0002-9947-1953-0054175-2.pdf). These proofs are in essence very similar, but the latter is more aesthetically pleasing, and it is this proof that appears in all the textbooks. I went looking and did not see another proof anywhere in the literature. We provide a few convenience constructors for `star_ordered_ring` in the form of reducible definitions which can apply when `R` is either a `non_unital_ring` (so we only need to characterize nonnegativity), and / or when positive elements have exactly the form `star s * s`. In this way, we can effectively maintain the status quo (see the instances for `real` and `complex`). --- docs/references.bib | 12 ++ src/algebra/star/basic.lean | 67 ------- src/algebra/star/chsh.lean | 6 +- src/algebra/star/order.lean | 179 ++++++++++++++++++ .../star/continuous_functional_calculus.lean | 2 +- src/data/complex/basic.lean | 25 +-- src/data/rat/star.lean | 50 +++++ src/data/real/sqrt.lean | 12 +- src/linear_algebra/matrix/dot_product.lean | 1 + 9 files changed, 266 insertions(+), 88 deletions(-) create mode 100644 src/algebra/star/order.lean create mode 100644 src/data/rat/star.lean diff --git a/docs/references.bib b/docs/references.bib index 51d5318f01f37..2569176a1a9f9 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -1473,6 +1473,18 @@ @Book{ kechris1995 url = {https://doi.org/10.1007/978-1-4612-4190-4} } +@Article{ kelleyVaught1953, + author = {Kelley, J. L. and Vaught, R. L.}, + title = {The positive cone in {Banach} algebras}, + journal = {Trans. Am. Math. Soc.}, + issn = {0002-9947}, + volume = {74}, + pages = {44--55}, + year = {1953}, + language = {English}, + doi = {10.2307/1990847} +} + @Article{ kleiman1979, author = {Kleiman, Steven Lawrence}, title = {Misconceptions about {$K\_X$}}, diff --git a/src/algebra/star/basic.lean b/src/algebra/star/basic.lean index 9b18517de2d79..04fb6397d361f 100644 --- a/src/algebra/star/basic.lean +++ b/src/algebra/star/basic.lean @@ -22,23 +22,12 @@ These are implemented as "mixin" typeclasses, so to summon a star ring (for exam one needs to write `(R : Type) [ring R] [star_ring R]`. This avoids difficulties with diamond inheritance. -We also define the class `star_ordered_ring R`, which says that the order on `R` respects the -star operation, i.e. an element `r` is nonnegative iff there exists an `s` such that -`r = star s * s`. - For now we simply do not introduce notations, as different users are expected to feel strongly about the relative merits of `r^*`, `r†`, `rᘁ`, and so on. Our star rings are actually star semirings, but of course we can prove `star_neg : star (-r) = - star r` when the underlying semiring is a ring. - -## TODO - -* In a Banach star algebra without a well-defined square root, the natural ordering is given by the -positive cone which is the closure of the sums of elements `star r * r`. A weaker version of -`star_ordered_ring` could be defined for this case. Note that the current definition has the -advantage of not requiring a topology. -/ assert_not_exists finset @@ -365,62 +354,6 @@ def star_ring_of_comm {R : Type*} [comm_semiring R] : star_ring R := star_add := λ x y, rfl, ..star_semigroup_of_comm } -/-- -An ordered `*`-ring is a ring which is both an `ordered_add_comm_group` and a `*`-ring, -and `0 ≤ r ↔ ∃ s, r = star s * s`. --/ -class star_ordered_ring (R : Type u) [non_unital_semiring R] [partial_order R] - extends star_ring R := -(add_le_add_left : ∀ a b : R, a ≤ b → ∀ c : R, c + a ≤ c + b) -(nonneg_iff : ∀ r : R, 0 ≤ r ↔ ∃ s, r = star s * s) - -namespace star_ordered_ring - -@[priority 100] -- see note [lower instance priority] -instance [non_unital_ring R] [partial_order R] [star_ordered_ring R] : ordered_add_comm_group R := -{ ..show non_unital_ring R, by apply_instance, - ..show partial_order R, by apply_instance, - ..show star_ordered_ring R, by apply_instance } - -end star_ordered_ring - -section non_unital_semiring - -variables [non_unital_semiring R] [partial_order R] [star_ordered_ring R] - -lemma star_mul_self_nonneg {r : R} : 0 ≤ star r * r := -(star_ordered_ring.nonneg_iff _).mpr ⟨r, rfl⟩ - -lemma star_mul_self_nonneg' {r : R} : 0 ≤ r * star r := -by { nth_rewrite_rhs 0 [←star_star r], exact star_mul_self_nonneg } - -lemma conjugate_nonneg {a : R} (ha : 0 ≤ a) (c : R) : 0 ≤ star c * a * c := -begin - obtain ⟨x, rfl⟩ := (star_ordered_ring.nonneg_iff _).1 ha, - exact (star_ordered_ring.nonneg_iff _).2 ⟨x * c, by rw [star_mul, ←mul_assoc, mul_assoc _ _ c]⟩, -end - -lemma conjugate_nonneg' {a : R} (ha : 0 ≤ a) (c : R) : 0 ≤ c * a * star c := -by simpa only [star_star] using conjugate_nonneg ha (star c) - -end non_unital_semiring - -section non_unital_ring - -variables [non_unital_ring R] [partial_order R] [star_ordered_ring R] - -lemma conjugate_le_conjugate {a b : R} (hab : a ≤ b) (c : R) : star c * a * c ≤ star c * b * c := -begin - rw ←sub_nonneg at hab ⊢, - convert conjugate_nonneg hab c, - simp only [mul_sub, sub_mul], -end - -lemma conjugate_le_conjugate' {a b : R} (hab : a ≤ b) (c : R) : c * a * star c ≤ c * b * star c := -by simpa only [star_star] using conjugate_le_conjugate hab (star c) - -end non_unital_ring - /-- A star module `A` over a star ring `R` is a module which is a star add monoid, and the two star structures are compatible in the sense diff --git a/src/algebra/star/chsh.lean b/src/algebra/star/chsh.lean index f794817618872..61e87cd570dca 100644 --- a/src/algebra/star/chsh.lean +++ b/src/algebra/star/chsh.lean @@ -133,7 +133,7 @@ begin T.A₀_sa, T.A₁_sa, T.B₀_sa, T.B₁_sa, mul_comm B₀, mul_comm B₁], }, rw idem', conv_rhs { congr, skip, congr, rw ←sa, }, - convert smul_le_smul_of_nonneg (star_mul_self_nonneg : 0 ≤ star P * P) _, + convert smul_le_smul_of_nonneg (star_mul_self_nonneg P) _, { simp, }, { apply_instance, }, { norm_num, } }, @@ -221,11 +221,11 @@ begin have P2_nonneg : 0 ≤ P^2, { rw [sq], conv { congr, skip, congr, rw ←P_sa, }, - convert (star_mul_self_nonneg : 0 ≤ star P * P), }, + convert (star_mul_self_nonneg P), }, have Q2_nonneg : 0 ≤ Q^2, { rw [sq], conv { congr, skip, congr, rw ←Q_sa, }, - convert (star_mul_self_nonneg : 0 ≤ star Q * Q), }, + convert (star_mul_self_nonneg Q), }, convert smul_le_smul_of_nonneg (add_nonneg P2_nonneg Q2_nonneg) (le_of_lt (show 0 < √2⁻¹, by norm_num)), -- `norm_num` can't directly show `0 ≤ √2⁻¹` simp, }, diff --git a/src/algebra/star/order.lean b/src/algebra/star/order.lean new file mode 100644 index 0000000000000..540a7acd6351f --- /dev/null +++ b/src/algebra/star/order.lean @@ -0,0 +1,179 @@ +/- +Copyright (c) 2023 Scott Morrison. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Scott Morrison +-/ + +import algebra.star.basic +import group_theory.submonoid.basic + +/-! # Star ordered rings + +We define the class `star_ordered_ring R`, which says that the order on `R` respects the +star operation, i.e. an element `r` is nonnegative iff it is in the `add_submonoid` generated by +elements of the form `star s * s`. In many cases, including all C⋆-algebras, this can be reduced to +`0 ≤ r ↔ ∃ s, r = star s * s`. However, this generality is slightly more convenient (e.g., it +allows us to register a `star_ordered_ring` instance for `ℚ`), and more closely resembles the +literature (see the seminal paper [*The positive cone in Banach algebras*][kelleyVaught1953]) + +In order to accodomate `non_unital_semiring R`, we actually don't characterize nonnegativity, but +rather the entire `≤` relation with `star_ordered_ring.le_iff`. However, notice that when `R` is a +`non_unital_ring`, these are equivalent (see `star_ordered_ring.nonneg_iff` and +`star_ordered_ring.of_nonneg_iff`). + +## TODO + +* In a Banach star algebra without a well-defined square root, the natural ordering is given by the +positive cone which is the _closure_ of the sums of elements `star r * r`. A weaker version of +`star_ordered_ring` could be defined for this case (again, see +[*The positive cone in Banach algebras*][kelleyVaught1953]). Note that the current definition has +the advantage of not requiring a topology. +-/ + +universe u +variable {R : Type u} + +/-- +An ordered `*`-ring is a ring which is both an `ordered_add_comm_group` and a `*`-ring, +and the nonnegative elements constitute precisely the `add_submonoid` generated by +elements of the form `star s * s`. + +If you are working with a `non_unital_ring` and not a `non_unital_semiring`, it may be more +convenient do declare instances using `star_ordered_ring.of_nonneg_iff'`. -/ +class star_ordered_ring (R : Type u) [non_unital_semiring R] [partial_order R] + extends star_ring R := +(add_le_add_left : ∀ a b : R, a ≤ b → ∀ c : R, c + a ≤ c + b) +(le_iff : ∀ x y : R, + x ≤ y ↔ ∃ p, p ∈ add_submonoid.closure (set.range $ λ s, star s * s) ∧ y = x + p) + +namespace star_ordered_ring + +@[priority 100] -- see note [lower instance priority] +instance to_ordered_add_comm_monoid [non_unital_semiring R] [partial_order R] + [star_ordered_ring R] : ordered_add_comm_monoid R := +{ ..show non_unital_semiring R, by apply_instance, + ..show partial_order R, by apply_instance, + ..show star_ordered_ring R, by apply_instance } + +@[priority 100] -- see note [lower instance priority] +instance to_has_exists_add_of_le [non_unital_semiring R] [partial_order R] + [star_ordered_ring R] : has_exists_add_of_le R := +{ exists_add_of_le := λ a b h, match (le_iff _ _).mp h with ⟨p, _, hp⟩ := ⟨p, hp⟩ end } + +@[priority 100] -- see note [lower instance priority] +instance to_ordered_add_comm_group [non_unital_ring R] [partial_order R] [star_ordered_ring R] : + ordered_add_comm_group R := +{ ..show non_unital_ring R, by apply_instance, + ..show partial_order R, by apply_instance, + ..show star_ordered_ring R, by apply_instance } + +/-- To construct a `star_ordered_ring` instance it suffices to show that `x ≤ y` if and only if +`y = x + star s * s` for some `s : R`. + +This is provided for convenience because it holds in some common scenarios (e.g.,`ℝ≥0`, `C(X, ℝ≥0)`) +and obviates the hassle of `add_submonoid.closure_induction` when creating those instances. + +If you are working with a `non_unital_ring` and not a `non_unital_semiring`, see +`star_ordered_ring.of_nonneg_iff` for a more convenient version. -/ +@[reducible] -- set note [reducible non-instances] +def of_le_iff [non_unital_semiring R] [partial_order R] [star_ring R] + (h_add : ∀ {x y : R}, x ≤ y → ∀ z, z + x ≤ z + y) + (h_le_iff : ∀ x y : R, x ≤ y ↔ ∃ s, y = x + star s * s) : + star_ordered_ring R := +{ add_le_add_left := @h_add, + le_iff := λ x y, + begin + refine ⟨λ h, _, _⟩, + { obtain ⟨p, hp⟩ := (h_le_iff x y).mp h, + exact ⟨star p * p, add_submonoid.subset_closure ⟨p, rfl⟩, hp⟩ }, + { rintro ⟨p, hp, hpxy⟩, + revert x y hpxy, + refine add_submonoid.closure_induction hp _ (λ x y h, add_zero x ▸ h.ge) _, + { rintro _ ⟨s, rfl⟩ x y rfl, + nth_rewrite 0 [←add_zero x], + refine h_add _ x, + exact (h_le_iff _ _).mpr ⟨s, by rw [zero_add]⟩ }, + { rintro a b ha hb x y rfl, + nth_rewrite 0 [←add_zero x], + refine h_add ((ha 0 _ (zero_add a).symm).trans (hb a _ rfl)) x } } + end, + .. ‹star_ring R› } + +/-- When `R` is a non-unital ring, to construct a `star_ordered_ring` instance it suffices to +show that the nonnegative elements are precisely those elements in the `add_submonoid` generated +by `star s * s` for `s : R`. -/ +@[reducible] -- set note [reducible non-instances] +def of_nonneg_iff [non_unital_ring R] [partial_order R] [star_ring R] + (h_add : ∀ {x y : R}, x ≤ y → ∀ z, z + x ≤ z + y) + (h_nonneg_iff : ∀ x : R, 0 ≤ x ↔ x ∈ add_submonoid.closure (set.range $ λ s : R, star s * s)) : + star_ordered_ring R := +{ add_le_add_left := @h_add, + le_iff := λ x y, + begin + haveI : covariant_class R R (+) (≤) := ⟨λ _ _ _ h, h_add h _⟩, + simpa only [←sub_eq_iff_eq_add', sub_nonneg, exists_eq_right'] using h_nonneg_iff (y - x), + end, + .. ‹star_ring R› } + +/-- When `R` is a non-unital ring, to construct a `star_ordered_ring` instance it suffices to +show that the nonnegative elements are precisely those elements of the form `star s * s` +for `s : R`. + +This is provided for convenience because it holds in many common scenarios (e.g.,`ℝ`, `ℂ`, or +any C⋆-algebra), and obviates the hassle of `add_submonoid.closure_induction` when creating those +instances. -/ +@[reducible] -- set note [reducible non-instances] +def of_nonneg_iff' [non_unital_ring R] [partial_order R] [star_ring R] + (h_add : ∀ {x y : R}, x ≤ y → ∀ z, z + x ≤ z + y) + (h_nonneg_iff : ∀ x : R, 0 ≤ x ↔ ∃ s, x = star s * s) : + star_ordered_ring R := +of_le_iff @h_add +begin + haveI : covariant_class R R (+) (≤) := ⟨λ _ _ _ h, h_add h _⟩, + simpa [sub_eq_iff_eq_add', sub_nonneg] using λ x y, h_nonneg_iff (y - x), +end + +lemma nonneg_iff [non_unital_semiring R] [partial_order R] [star_ordered_ring R] + {x : R} : 0 ≤ x ↔ x ∈ add_submonoid.closure (set.range $ λ s : R, star s * s) := +by simp only [le_iff, zero_add, exists_eq_right'] + +end star_ordered_ring + +section non_unital_semiring + +variables [non_unital_semiring R] [partial_order R] [star_ordered_ring R] + +lemma star_mul_self_nonneg (r : R) : 0 ≤ star r * r := +star_ordered_ring.nonneg_iff.mpr $ add_submonoid.subset_closure ⟨r, rfl⟩ + +lemma star_mul_self_nonneg' (r : R) : 0 ≤ r * star r := +by { nth_rewrite_rhs 0 [←star_star r], exact star_mul_self_nonneg (star r) } + +lemma conjugate_nonneg {a : R} (ha : 0 ≤ a) (c : R) : 0 ≤ star c * a * c := +begin + rw star_ordered_ring.nonneg_iff at ha, + refine add_submonoid.closure_induction ha (λ x hx, _) (by rw [mul_zero, zero_mul]) + (λ x y hx hy, _), + { obtain ⟨x, rfl⟩ := hx, + convert star_mul_self_nonneg (x * c) using 1, + rw [star_mul, ←mul_assoc, mul_assoc _ _ c] }, + { calc 0 ≤ star c * x * c + 0 : by rw [add_zero]; exact hx + ... ≤ star c * x * c + star c * y * c : star_ordered_ring.add_le_add_left _ _ hy _ + ... ≤ _ : by rw [mul_add, add_mul] } +end + +lemma conjugate_nonneg' {a : R} (ha : 0 ≤ a) (c : R) : 0 ≤ c * a * star c := +by simpa only [star_star] using conjugate_nonneg ha (star c) + +lemma conjugate_le_conjugate {a b : R} (hab : a ≤ b) (c : R) : star c * a * c ≤ star c * b * c := +begin + rw [star_ordered_ring.le_iff] at hab ⊢, + obtain ⟨p, hp, rfl⟩ := hab, + simp_rw [←star_ordered_ring.nonneg_iff] at hp ⊢, + exact ⟨star c * p * c, conjugate_nonneg hp c, by simp only [add_mul, mul_add]⟩, +end + +lemma conjugate_le_conjugate' {a b : R} (hab : a ≤ b) (c : R) : c * a * star c ≤ c * b * star c := +by simpa only [star_star] using conjugate_le_conjugate hab (star c) + +end non_unital_semiring diff --git a/src/analysis/normed_space/star/continuous_functional_calculus.lean b/src/analysis/normed_space/star/continuous_functional_calculus.lean index df35d3e3b9eb6..f3abf11cc57f4 100644 --- a/src/analysis/normed_space/star/continuous_functional_calculus.lean +++ b/src/analysis/normed_space/star/continuous_functional_calculus.lean @@ -100,7 +100,7 @@ begin rw [←spectrum.gelfand_transform_eq (star a' * a'), continuous_map.spectrum_eq_range], rintro - ⟨φ, rfl⟩, rw [gelfand_transform_apply_apply ℂ _ (star a' * a') φ, map_mul φ, map_star φ], - rw [complex.eq_coe_norm_of_nonneg star_mul_self_nonneg, ←map_star, ←map_mul], + rw [complex.eq_coe_norm_of_nonneg (star_mul_self_nonneg _), ←map_star, ←map_mul], exact ⟨complex.zero_le_real.2 (norm_nonneg _), complex.real_le_real.2 (alg_hom.norm_apply_le_self φ (star a' * a'))⟩, } end diff --git a/src/data/complex/basic.lean b/src/data/complex/basic.lean index 74951e27e6bbf..9644405e9b04f 100644 --- a/src/data/complex/basic.lean +++ b/src/data/complex/basic.lean @@ -666,18 +666,19 @@ With `z ≤ w` iff `w - z` is real and nonnegative, `ℂ` is a star ordered ring (That is, a star ring in which the nonnegative elements are those of the form `star z * z`.) -/ protected def star_ordered_ring : star_ordered_ring ℂ := -{ nonneg_iff := λ r, by - { refine ⟨λ hr, ⟨real.sqrt r.re, _⟩, λ h, _⟩, - { have h₁ : 0 ≤ r.re := by { rw [le_def] at hr, exact hr.1 }, - have h₂ : r.im = 0 := by { rw [le_def] at hr, exact hr.2.symm }, - ext, - { simp only [of_real_im, star_def, of_real_re, sub_zero, conj_re, mul_re, mul_zero, - ←real.sqrt_mul h₁ r.re, real.sqrt_mul_self h₁] }, - { simp only [h₂, add_zero, of_real_im, star_def, zero_mul, conj_im, - mul_im, mul_zero, neg_zero] } }, - { obtain ⟨s, rfl⟩ := h, - simp only [←norm_sq_eq_conj_mul_self, norm_sq_nonneg, zero_le_real, star_def] } }, - ..complex.strict_ordered_comm_ring } +star_ordered_ring.of_nonneg_iff' (λ _ _, add_le_add_left) $ λ r, +begin + refine ⟨λ hr, ⟨real.sqrt r.re, _⟩, λ h, _⟩, + { have h₁ : 0 ≤ r.re := by { rw [le_def] at hr, exact hr.1 }, + have h₂ : r.im = 0 := by { rw [le_def] at hr, exact hr.2.symm }, + ext, + { simp only [of_real_im, star_def, of_real_re, sub_zero, conj_re, mul_re, mul_zero, + ←real.sqrt_mul h₁ r.re, real.sqrt_mul_self h₁] }, + { simp only [h₂, add_zero, of_real_im, star_def, zero_mul, conj_im, + mul_im, mul_zero, neg_zero] } }, + { obtain ⟨s, rfl⟩ := h, + simp only [←norm_sq_eq_conj_mul_self, norm_sq_nonneg, zero_le_real, star_def] }, +end localized "attribute [instance] complex.star_ordered_ring" in complex_order diff --git a/src/data/rat/star.lean b/src/data/rat/star.lean new file mode 100644 index 0000000000000..e254482fe566e --- /dev/null +++ b/src/data/rat/star.lean @@ -0,0 +1,50 @@ +/- +Copyright (c) 2023 Jireh Loreaux. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jireh Loreaux +-/ + +import algebra.star.order +import data.rat.lemmas +import group_theory.submonoid.membership + +/-! # Star order structure on ℚ + +Here we put the trivial `star` operation on `ℚ` for convenience and show that it is a +`star_ordered_ring`. In particular, this means that every element of `ℚ` is a sum of squares. +-/ + +namespace rat + +instance : star_ring ℚ := +{ star := id, + star_involutive := λ _, rfl, + star_mul := λ _ _, mul_comm _ _, + star_add := λ _ _, rfl } + +instance : has_trivial_star ℚ := +{ star_trivial := λ _, rfl } + +instance : star_ordered_ring ℚ := +star_ordered_ring.of_nonneg_iff (λ _ _, add_le_add_left) $ λ x, +begin + refine ⟨λ hx, _, λ hx, add_submonoid.closure_induction hx + (by { rintro - ⟨s, rfl⟩, exact mul_self_nonneg s }) le_rfl (λ _ _, add_nonneg)⟩, + /- If `x = p / q`, then, since `0 ≤ x`, we have `p q : ℕ`, and `p / q` is the sum of `p * q` + copies of `(1 / q) ^ 2`, and so `x` lies in the `add_submonoid` generated by square elements. + + Note: it's possible to rephrase this argument as `x = (p * q) • (1 / q) ^ 2`, but this would + be somewhat challenging without increasing import requirements. -/ + suffices : (finset.range (x.num.nat_abs * x.denom)).sum + (function.const ℕ (rat.mk_pnat 1 ⟨x.denom, x.pos⟩ * rat.mk_pnat 1 ⟨x.denom, x.pos⟩)) = x, + { exact this ▸ sum_mem (λ n hn, add_submonoid.subset_closure ⟨_, rfl⟩) }, + simp only [function.const_apply, finset.sum_const, finset.card_range, nsmul_eq_mul, mk_pnat_eq], + rw [←int.cast_coe_nat, int.coe_nat_mul, int.coe_nat_abs, + abs_of_nonneg (num_nonneg_iff_zero_le.mpr hx), int.cast_mul, int.cast_coe_nat], + simp only [int.cast_mul, int.cast_coe_nat, coe_int_eq_mk, coe_nat_eq_mk], + rw [mul_assoc, ←mul_assoc (mk (x.denom : ℤ) 1), mk_mul_mk_cancel one_ne_zero, + ←one_mul (x.denom : ℤ), div_mk_div_cancel_left (by simpa using x.pos.ne' : (x.denom : ℤ) ≠ 0), + one_mul, mk_one_one, one_mul, mk_mul_mk_cancel one_ne_zero, rat.num_denom], +end + +end rat diff --git a/src/data/real/sqrt.lean b/src/data/real/sqrt.lean index c8f1337bd8715..c3ae85a902843 100644 --- a/src/data/real/sqrt.lean +++ b/src/data/real/sqrt.lean @@ -3,6 +3,7 @@ Copyright (c) 2020 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Floris van Doorn, Yury Kudryashov -/ +import algebra.star.order import topology.algebra.order.monotone_continuity import topology.instances.nnreal import tactic.positivity @@ -340,11 +341,12 @@ begin end instance : star_ordered_ring ℝ := -{ nonneg_iff := λ r, by - { refine ⟨λ hr, ⟨sqrt r, show r = sqrt r * sqrt r, by rw [←sqrt_mul hr, sqrt_mul_self hr]⟩, _⟩, - rintros ⟨s, rfl⟩, - exact mul_self_nonneg s }, - ..real.ordered_add_comm_group } +star_ordered_ring.of_nonneg_iff' (λ _ _, add_le_add_left) $ λ r, +begin + refine ⟨λ hr, ⟨sqrt r, show r = sqrt r * sqrt r, by rw [←sqrt_mul hr, sqrt_mul_self hr]⟩, _⟩, + rintros ⟨s, rfl⟩, + exact mul_self_nonneg s +end end real diff --git a/src/linear_algebra/matrix/dot_product.lean b/src/linear_algebra/matrix/dot_product.lean index 1aca66ab84fd6..e3c81c8f65d16 100644 --- a/src/linear_algebra/matrix/dot_product.lean +++ b/src/linear_algebra/matrix/dot_product.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Patrick Massot, Casper Putz, Anne Baanen -/ +import algebra.star.order import data.matrix.basic import linear_algebra.std_basis From b915e9392ecb2a861e1e766f0e1df6ac481188ca Mon Sep 17 00:00:00 2001 From: Eric Rodriguez Date: Wed, 7 Jun 2023 19:34:28 +0000 Subject: [PATCH 1157/1291] feat(ring_theory/subring): centralizer as a subring (#18861) Co-authored-by: Eric Rodriguez <37984851+ericrbg@users.noreply.github.com> --- src/algebra/algebra/subalgebra/basic.lean | 5 +++ src/group_theory/subgroup/basic.lean | 5 +++ src/group_theory/submonoid/centralizer.lean | 5 +++ .../subsemigroup/centralizer.lean | 14 ++++++++ .../non_unital_subsemiring/basic.lean | 7 ++++ src/ring_theory/subring/basic.lean | 33 +++++++++++++++++++ src/ring_theory/subsemiring/basic.lean | 7 ++++ 7 files changed, 76 insertions(+) diff --git a/src/algebra/algebra/subalgebra/basic.lean b/src/algebra/algebra/subalgebra/basic.lean index 9cc6730b0237d..f0ef0af5aec09 100644 --- a/src/algebra/algebra/subalgebra/basic.lean +++ b/src/algebra/algebra/subalgebra/basic.lean @@ -1049,10 +1049,15 @@ lemma mem_centralizer_iff {s : set A} {z : A} : z ∈ centralizer R s ↔ ∀ g ∈ s, g * z = z * g := iff.rfl +lemma center_le_centralizer (s) : center R A ≤ centralizer R s := s.center_subset_centralizer + lemma centralizer_le (s t : set A) (h : s ⊆ t) : centralizer R t ≤ centralizer R s := set.centralizer_subset h +@[simp] lemma centralizer_eq_top_iff_subset {s : set A} : centralizer R s = ⊤ ↔ s ⊆ center R A := +set_like.ext'_iff.trans set.centralizer_eq_top_iff_subset + @[simp] lemma centralizer_univ : centralizer R set.univ = center R A := set_like.ext' (set.centralizer_univ A) diff --git a/src/group_theory/subgroup/basic.lean b/src/group_theory/subgroup/basic.lean index b18053c47065b..706259668852c 100644 --- a/src/group_theory/subgroup/basic.lean +++ b/src/group_theory/subgroup/basic.lean @@ -1604,9 +1604,14 @@ set_like.ext' (set.centralizer_univ G) @[to_additive] lemma le_centralizer_iff : H ≤ K.centralizer ↔ K ≤ H.centralizer := ⟨λ h x hx y hy, (h hy x hx).symm, λ h x hx y hy, (h hy x hx).symm⟩ +lemma center_le_centralizer (s) : center G ≤ centralizer s := set.center_subset_centralizer s + @[to_additive] lemma centralizer_le (h : H ≤ K) : centralizer K ≤ centralizer H := submonoid.centralizer_le h +@[simp] lemma centralizer_eq_top_iff_subset {s} : centralizer s = ⊤ ↔ s ≤ center G := +set_like.ext'_iff.trans set.centralizer_eq_top_iff_subset + @[to_additive] instance subgroup.centralizer.characteristic [hH : H.characteristic] : H.centralizer.characteristic := begin diff --git a/src/group_theory/submonoid/centralizer.lean b/src/group_theory/submonoid/centralizer.lean index f05395a009469..47bc7785fb91c 100644 --- a/src/group_theory/submonoid/centralizer.lean +++ b/src/group_theory/submonoid/centralizer.lean @@ -50,6 +50,8 @@ variables {S} @[to_additive] lemma mem_centralizer_iff {z : M} : z ∈ centralizer S ↔ ∀ g ∈ S, g * z = z * g := iff.rfl +lemma center_le_centralizer (s) : center M ≤ centralizer s := s.center_subset_centralizer + @[to_additive] instance decidable_mem_centralizer (a) [decidable $ ∀ b ∈ S, b * a = a * b] : decidable (a ∈ centralizer S) := decidable_of_iff' _ mem_centralizer_iff @@ -58,6 +60,9 @@ decidable_of_iff' _ mem_centralizer_iff lemma centralizer_le (h : S ⊆ T) : centralizer T ≤ centralizer S := set.centralizer_subset h +@[simp] lemma centralizer_eq_top_iff_subset {s : set M} : centralizer s = ⊤ ↔ s ⊆ center M := +set_like.ext'_iff.trans set.centralizer_eq_top_iff_subset + variables (M) @[simp, to_additive] diff --git a/src/group_theory/subsemigroup/centralizer.lean b/src/group_theory/subsemigroup/centralizer.lean index 8eadf8a074db3..e7dda9535bea3 100644 --- a/src/group_theory/subsemigroup/centralizer.lean +++ b/src/group_theory/subsemigroup/centralizer.lean @@ -100,6 +100,14 @@ end lemma centralizer_subset [has_mul M] (h : S ⊆ T) : centralizer T ⊆ centralizer S := λ t ht s hs, ht s (h hs) +lemma center_subset_centralizer [has_mul M] (S : set M) : set.center M ⊆ S.centralizer := +λ x hx m _, hx m + +@[simp] lemma centralizer_eq_top_iff_subset {s : set M} [has_mul M] : + centralizer s = set.univ ↔ s ⊆ center M := +eq_top_iff.trans $ ⟨λ h x hx g, (h trivial _ hx).symm, + λ h x _ m hm, (h hm x).symm⟩ + variables (M) @[simp, to_additive add_centralizer_univ] @@ -112,6 +120,7 @@ variables {M} (S) 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 subsemigroup @@ -135,10 +144,15 @@ iff.rfl decidable (a ∈ centralizer S) := decidable_of_iff' _ mem_centralizer_iff +lemma center_le_centralizer (S) : center M ≤ centralizer S := S.center_subset_centralizer + @[to_additive] lemma centralizer_le (h : S ⊆ T) : centralizer T ≤ centralizer S := set.centralizer_subset h +@[simp] lemma centralizer_eq_top_iff_subset {s : set M} : centralizer s = ⊤ ↔ s ⊆ center M := +set_like.ext'_iff.trans set.centralizer_eq_top_iff_subset + variables (M) @[simp, to_additive] diff --git a/src/ring_theory/non_unital_subsemiring/basic.lean b/src/ring_theory/non_unital_subsemiring/basic.lean index 2368f132629a9..8b32d68f38039 100644 --- a/src/ring_theory/non_unital_subsemiring/basic.lean +++ b/src/ring_theory/non_unital_subsemiring/basic.lean @@ -417,10 +417,17 @@ lemma mem_centralizer_iff {R} [non_unital_semiring R] {s : set R} {z : R} : z ∈ centralizer s ↔ ∀ g ∈ s, g * z = z * g := iff.rfl +lemma center_le_centralizer {R} [non_unital_semiring R] (s) : center R ≤ centralizer s := + s.center_subset_centralizer + lemma centralizer_le {R} [non_unital_semiring R] (s t : set R) (h : s ⊆ t) : centralizer t ≤ centralizer s := set.centralizer_subset h +@[simp] lemma centralizer_eq_top_iff_subset {R} [non_unital_semiring R] {s : set R} : + centralizer s = ⊤ ↔ s ⊆ center R := +set_like.ext'_iff.trans set.centralizer_eq_top_iff_subset + @[simp] lemma centralizer_univ {R} [non_unital_semiring R] : centralizer set.univ = center R := set_like.ext' (set.centralizer_univ R) diff --git a/src/ring_theory/subring/basic.lean b/src/ring_theory/subring/basic.lean index 68334a8fdd49d..4df02ed3ade6b 100644 --- a/src/ring_theory/subring/basic.lean +++ b/src/ring_theory/subring/basic.lean @@ -651,6 +651,39 @@ lemma center.coe_div (a b : center K) : ((a / b : center K) : K) = (a : K) / (b end division_ring +section centralizer + +/-- The centralizer of a set inside a ring as a `subring`. -/ +def centralizer (s : set R) : subring R := +{ neg_mem' := λ x, set.neg_mem_centralizer, + ..subsemiring.centralizer s } + +@[simp, norm_cast] +lemma coe_centralizer (s : set R) : (centralizer s : set R) = s.centralizer := rfl + +lemma centralizer_to_submonoid (s : set R) : + (centralizer s).to_submonoid = submonoid.centralizer s := rfl + +lemma centralizer_to_subsemiring (s : set R) : + (centralizer s).to_subsemiring = subsemiring.centralizer s := rfl + +lemma mem_centralizer_iff {s : set R} {z : R} : + z ∈ centralizer s ↔ ∀ g ∈ s, g * z = z * g := +iff.rfl + +lemma center_le_centralizer (s) : center R ≤ centralizer s := s.center_subset_centralizer + +lemma centralizer_le (s t : set R) (h : s ⊆ t) : centralizer t ≤ centralizer s := +set.centralizer_subset h + +@[simp] lemma centralizer_eq_top_iff_subset {s : set R} : centralizer s = ⊤ ↔ s ⊆ center R := +set_like.ext'_iff.trans set.centralizer_eq_top_iff_subset + +@[simp] lemma centralizer_univ : centralizer set.univ = center R := +set_like.ext' (set.centralizer_univ R) + +end centralizer + /-! ## subring closure of a subset -/ /-- The `subring` generated by a set. -/ diff --git a/src/ring_theory/subsemiring/basic.lean b/src/ring_theory/subsemiring/basic.lean index f06b9c51eff95..4794203c5798c 100644 --- a/src/ring_theory/subsemiring/basic.lean +++ b/src/ring_theory/subsemiring/basic.lean @@ -604,10 +604,17 @@ lemma mem_centralizer_iff {R} [semiring R] {s : set R} {z : R} : z ∈ centralizer s ↔ ∀ g ∈ s, g * z = z * g := iff.rfl +lemma center_le_centralizer {R} [semiring R] (s) : center R ≤ centralizer s := + s.center_subset_centralizer + lemma centralizer_le {R} [semiring R] (s t : set R) (h : s ⊆ t) : centralizer t ≤ centralizer s := set.centralizer_subset h +@[simp] lemma centralizer_eq_top_iff_subset {R} [semiring R] {s : set R} : + centralizer s = ⊤ ↔ s ⊆ center R := +set_like.ext'_iff.trans set.centralizer_eq_top_iff_subset + @[simp] lemma centralizer_univ {R} [semiring R] : centralizer set.univ = center R := set_like.ext' (set.centralizer_univ R) From 7e5137f579de09a059a5ce98f364a04e221aabf0 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Thu, 8 Jun 2023 04:42:31 +0000 Subject: [PATCH 1158/1291] chore(*): add mathlib4 synchronization comments (#19165) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.lie.classical` * `analysis.calculus.bump_function_inner` * `analysis.calculus.monotone` * `analysis.calculus.parametric_interval_integral` * `analysis.normed_space.quaternion_exponential` * `analysis.special_functions.trigonometric.arctan_deriv` * `analysis.special_functions.trigonometric.bounds` * `category_theory.bicategory.coherence_tactic` * `category_theory.monoidal.Mon_` * `category_theory.monoidal.coherence` * `category_theory.monoidal.subcategory` * `category_theory.sites.compatible_plus` * `category_theory.sites.types` * `data.json` * `data.real.pi.leibniz` * `geometry.euclidean.angle.unoriented.right_angle` * `linear_algebra.tensor_algebra.basic` * `measure_theory.covering.density_theorem` * `measure_theory.covering.differentiation` * `measure_theory.covering.liminf_limsup` * `measure_theory.covering.one_dim` * `measure_theory.decomposition.lebesgue` * `measure_theory.decomposition.radon_nikodym` * `measure_theory.function.continuous_map_dense` * `measure_theory.function.l2_space` * `measure_theory.function.special_functions.arctan` * `measure_theory.integral.interval_average` * `measure_theory.integral.interval_integral` * `measure_theory.integral.peak_function` * `measure_theory.measure.finite_measure` * `measure_theory.measure.haar.quotient` * `measure_theory.measure.probability_measure` * `probability.integration` * `ring_theory.class_group` * `ring_theory.roots_of_unity.basic` * `ring_theory.roots_of_unity.complex` * `topology.order.hom.esakia` --- src/algebra/lie/classical.lean | 3 +++ src/analysis/calculus/bump_function_inner.lean | 3 +++ src/analysis/calculus/monotone.lean | 3 +++ src/analysis/calculus/parametric_interval_integral.lean | 3 +++ src/analysis/normed_space/quaternion_exponential.lean | 3 +++ .../special_functions/trigonometric/arctan_deriv.lean | 3 +++ src/analysis/special_functions/trigonometric/bounds.lean | 3 +++ src/category_theory/bicategory/coherence_tactic.lean | 3 +++ src/category_theory/monoidal/Mon_.lean | 3 +++ src/category_theory/monoidal/coherence.lean | 3 +++ src/category_theory/monoidal/subcategory.lean | 3 +++ src/category_theory/sites/compatible_plus.lean | 3 +++ src/category_theory/sites/types.lean | 3 +++ src/data/json.lean | 3 +++ src/data/real/pi/leibniz.lean | 5 ++++- src/geometry/euclidean/angle/unoriented/right_angle.lean | 3 +++ src/linear_algebra/tensor_algebra/basic.lean | 3 +++ src/measure_theory/covering/density_theorem.lean | 3 +++ src/measure_theory/covering/differentiation.lean | 3 +++ src/measure_theory/covering/liminf_limsup.lean | 3 +++ src/measure_theory/covering/one_dim.lean | 3 +++ src/measure_theory/decomposition/lebesgue.lean | 3 +++ src/measure_theory/decomposition/radon_nikodym.lean | 3 +++ src/measure_theory/function/continuous_map_dense.lean | 3 +++ src/measure_theory/function/l2_space.lean | 3 +++ src/measure_theory/function/special_functions/arctan.lean | 3 +++ src/measure_theory/integral/interval_average.lean | 3 +++ src/measure_theory/integral/interval_integral.lean | 3 +++ src/measure_theory/integral/peak_function.lean | 3 +++ src/measure_theory/measure/finite_measure.lean | 3 +++ src/measure_theory/measure/haar/quotient.lean | 3 +++ src/measure_theory/measure/probability_measure.lean | 3 +++ src/probability/integration.lean | 3 +++ src/ring_theory/class_group.lean | 3 +++ src/ring_theory/roots_of_unity/basic.lean | 3 +++ src/ring_theory/roots_of_unity/complex.lean | 3 +++ src/topology/order/hom/esakia.lean | 3 +++ 37 files changed, 112 insertions(+), 1 deletion(-) diff --git a/src/algebra/lie/classical.lean b/src/algebra/lie/classical.lean index 7513bfe024462..b6b6e38cfb15d 100644 --- a/src/algebra/lie/classical.lean +++ b/src/algebra/lie/classical.lean @@ -14,6 +14,9 @@ import linear_algebra.symplectic_group /-! # Classical Lie algebras +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file is the place to find definitions and basic properties of the classical Lie algebras: * Aₗ = sl(l+1) * Bₗ ≃ so(l+1, l) ≃ so(2l+1) diff --git a/src/analysis/calculus/bump_function_inner.lean b/src/analysis/calculus/bump_function_inner.lean index 4dcda6d14946a..699f2dbea150b 100644 --- a/src/analysis/calculus/bump_function_inner.lean +++ b/src/analysis/calculus/bump_function_inner.lean @@ -13,6 +13,9 @@ import measure_theory.integral.set_integral /-! # Infinitely smooth bump function +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we construct several infinitely smooth functions with properties that an analytic function cannot have: diff --git a/src/analysis/calculus/monotone.lean b/src/analysis/calculus/monotone.lean index d20ad597d8a53..2281735450d5e 100644 --- a/src/analysis/calculus/monotone.lean +++ b/src/analysis/calculus/monotone.lean @@ -10,6 +10,9 @@ import order.monotone.extension /-! # Differentiability of monotone functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We show that a monotone function `f : ℝ → ℝ` is differentiable almost everywhere, in `monotone.ae_differentiable_at`. (We also give a version for a function monotone on a set, in `monotone_on.ae_differentiable_within_at`.) diff --git a/src/analysis/calculus/parametric_interval_integral.lean b/src/analysis/calculus/parametric_interval_integral.lean index 807cd3f5df161..0d9494850c02e 100644 --- a/src/analysis/calculus/parametric_interval_integral.lean +++ b/src/analysis/calculus/parametric_interval_integral.lean @@ -9,6 +9,9 @@ import measure_theory.integral.interval_integral /-! # Derivatives of interval integrals depending on parameters +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we restate theorems about derivatives of integrals depending on parameters for interval integrals. -/ diff --git a/src/analysis/normed_space/quaternion_exponential.lean b/src/analysis/normed_space/quaternion_exponential.lean index 4a6786af9568d..a92e0d96e7dd7 100644 --- a/src/analysis/normed_space/quaternion_exponential.lean +++ b/src/analysis/normed_space/quaternion_exponential.lean @@ -10,6 +10,9 @@ import analysis.special_functions.trigonometric.series /-! # Lemmas about `exp` on `quaternion`s +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains results about `exp` on `quaternion ℝ`. ## Main results diff --git a/src/analysis/special_functions/trigonometric/arctan_deriv.lean b/src/analysis/special_functions/trigonometric/arctan_deriv.lean index d36c11cf4d688..b34a13876c806 100644 --- a/src/analysis/special_functions/trigonometric/arctan_deriv.lean +++ b/src/analysis/special_functions/trigonometric/arctan_deriv.lean @@ -9,6 +9,9 @@ import analysis.special_functions.trigonometric.complex_deriv /-! # Derivatives of the `tan` and `arctan` functions. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Continuity and derivatives of the tangent and arctangent functions. -/ diff --git a/src/analysis/special_functions/trigonometric/bounds.lean b/src/analysis/special_functions/trigonometric/bounds.lean index 7e2e8244e4e08..936a9e1e98434 100644 --- a/src/analysis/special_functions/trigonometric/bounds.lean +++ b/src/analysis/special_functions/trigonometric/bounds.lean @@ -8,6 +8,9 @@ import analysis.special_functions.trigonometric.arctan_deriv /-! # Polynomial bounds for trigonometric functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main statements This file contains upper and lower bounds for real trigonometric functions in terms diff --git a/src/category_theory/bicategory/coherence_tactic.lean b/src/category_theory/bicategory/coherence_tactic.lean index f999bdc5e4b95..b71ac18eb4264 100644 --- a/src/category_theory/bicategory/coherence_tactic.lean +++ b/src/category_theory/bicategory/coherence_tactic.lean @@ -8,6 +8,9 @@ import category_theory.bicategory.coherence /-! # A `coherence` tactic for bicategories, and `⊗≫` (composition up to associators) +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + 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 diff --git a/src/category_theory/monoidal/Mon_.lean b/src/category_theory/monoidal/Mon_.lean index 25f12394f3d14..5a26a67311fce 100644 --- a/src/category_theory/monoidal/Mon_.lean +++ b/src/category_theory/monoidal/Mon_.lean @@ -12,6 +12,9 @@ import algebra.punit_instances /-! # The category of monoids in a monoidal category. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define monoids in a monoidal category `C` and show that the category of monoids is equivalent to the category of lax monoidal functors from the unit monoidal category to `C`. We also show that if `C` is braided, then the category of monoids is naturally monoidal. diff --git a/src/category_theory/monoidal/coherence.lean b/src/category_theory/monoidal/coherence.lean index 27b7e7bc6b669..33035d50beba4 100644 --- a/src/category_theory/monoidal/coherence.lean +++ b/src/category_theory/monoidal/coherence.lean @@ -9,6 +9,9 @@ import category_theory.bicategory.coherence_tactic /-! # A `coherence` tactic for monoidal categories, and `⊗≫` (composition up to associators) +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We provide a `coherence` tactic, which proves equations where the two sides differ by replacing strings of monoidal structural morphisms with other such strings. diff --git a/src/category_theory/monoidal/subcategory.lean b/src/category_theory/monoidal/subcategory.lean index 19288fe79a209..a1d995d799d5b 100644 --- a/src/category_theory/monoidal/subcategory.lean +++ b/src/category_theory/monoidal/subcategory.lean @@ -12,6 +12,9 @@ import category_theory.closed.monoidal /-! # Full monoidal subcategories +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given a monidal category `C` and a monoidal predicate on `C`, that is a function `P : C → Prop` closed under `𝟙_` and `⊗`, we can put a monoidal structure on `{X : C // P X}` (the category structure is defined in `category_theory.full_subcategory`). diff --git a/src/category_theory/sites/compatible_plus.lean b/src/category_theory/sites/compatible_plus.lean index dc0f90947a658..75a5242fe1bf1 100644 --- a/src/category_theory/sites/compatible_plus.lean +++ b/src/category_theory/sites/compatible_plus.lean @@ -7,6 +7,9 @@ import category_theory.sites.whiskering import category_theory.sites.plus /-! +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we prove that the plus functor is compatible with functors which preserve the correct limits and colimits. diff --git a/src/category_theory/sites/types.lean b/src/category_theory/sites/types.lean index dcbea865d72c1..1c5d7b81c911c 100644 --- a/src/category_theory/sites/types.lean +++ b/src/category_theory/sites/types.lean @@ -9,6 +9,9 @@ import category_theory.sites.canonical /-! # Grothendieck Topology and Sheaves on the Category of Types +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define a Grothendieck topology on the category of types, and construct the canonical functor that sends a type to a sheaf over the category of types, and make this an equivalence of categories. diff --git a/src/data/json.lean b/src/data/json.lean index 5f13ebb66e7f0..b504379d36475 100644 --- a/src/data/json.lean +++ b/src/data/json.lean @@ -8,6 +8,9 @@ import tactic.core /-! # Json serialization typeclass +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file provides helpers for serializing primitive types to json. `@[derive non_null_json_serializable]` will make any structure json serializable; for instance, diff --git a/src/data/real/pi/leibniz.lean b/src/data/real/pi/leibniz.lean index be7ce817b362d..7d26a16f6aa2b 100644 --- a/src/data/real/pi/leibniz.lean +++ b/src/data/real/pi/leibniz.lean @@ -5,7 +5,10 @@ Authors: Benjamin Davidson -/ import analysis.special_functions.trigonometric.arctan_deriv -/-! ### Leibniz's Series for Pi -/ +/-! ### Leibniz's Series for Pi + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4.-/ namespace real diff --git a/src/geometry/euclidean/angle/unoriented/right_angle.lean b/src/geometry/euclidean/angle/unoriented/right_angle.lean index 75beb298c6074..b6f03ee3b3454 100644 --- a/src/geometry/euclidean/angle/unoriented/right_angle.lean +++ b/src/geometry/euclidean/angle/unoriented/right_angle.lean @@ -9,6 +9,9 @@ import geometry.euclidean.angle.unoriented.affine /-! # Right-angled triangles +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves basic geometrical results about distances and angles in (possibly degenerate) right-angled triangles in real inner product spaces and Euclidean affine spaces. diff --git a/src/linear_algebra/tensor_algebra/basic.lean b/src/linear_algebra/tensor_algebra/basic.lean index 98229da93f802..099b0a0344048 100644 --- a/src/linear_algebra/tensor_algebra/basic.lean +++ b/src/linear_algebra/tensor_algebra/basic.lean @@ -12,6 +12,9 @@ import linear_algebra.multilinear.basic /-! # Tensor Algebras +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given a commutative semiring `R`, and an `R`-module `M`, we construct the tensor algebra of `M`. This is the free `R`-algebra generated (`R`-linearly) by the module `M`. diff --git a/src/measure_theory/covering/density_theorem.lean b/src/measure_theory/covering/density_theorem.lean index b94da4ad6e87f..143854e2f2ca9 100644 --- a/src/measure_theory/covering/density_theorem.lean +++ b/src/measure_theory/covering/density_theorem.lean @@ -10,6 +10,9 @@ import measure_theory.covering.differentiation /-! # Uniformly locally doubling measures and Lebesgue's density theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Lebesgue's density theorem states that given a set `S` in a sigma compact metric space with locally-finite uniformly locally doubling measure `μ` then for almost all points `x` in `S`, for any sequence of closed balls `B₀, B₁, B₂, ...` containing `x`, the limit `μ (S ∩ Bⱼ) / μ (Bⱼ) → 1` as diff --git a/src/measure_theory/covering/differentiation.lean b/src/measure_theory/covering/differentiation.lean index 8dd8cd38b8c53..7cf9a8478abf0 100644 --- a/src/measure_theory/covering/differentiation.lean +++ b/src/measure_theory/covering/differentiation.lean @@ -13,6 +13,9 @@ import measure_theory.decomposition.lebesgue /-! # Differentiation of measures +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + On a second countable metric space with a measure `μ`, consider a Vitali family (i.e., for each `x` one has a family of sets shrinking to `x`, with a good behavior with respect to covering theorems). Consider also another measure `ρ`. Then, for almost every `x`, the ratio `ρ a / μ a` converges when diff --git a/src/measure_theory/covering/liminf_limsup.lean b/src/measure_theory/covering/liminf_limsup.lean index b32a298afbbd7..bd254f5dcaff9 100644 --- a/src/measure_theory/covering/liminf_limsup.lean +++ b/src/measure_theory/covering/liminf_limsup.lean @@ -8,6 +8,9 @@ import measure_theory.covering.density_theorem /-! # Liminf, limsup, and uniformly locally doubling measures. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file is a place to collect lemmas about liminf and limsup for subsets of a metric space carrying a uniformly locally doubling measure. diff --git a/src/measure_theory/covering/one_dim.lean b/src/measure_theory/covering/one_dim.lean index 305f17edf8f87..851314e77b0ca 100644 --- a/src/measure_theory/covering/one_dim.lean +++ b/src/measure_theory/covering/one_dim.lean @@ -9,6 +9,9 @@ import measure_theory.measure.lebesgue.eq_haar /-! # Covering theorems for Lebesgue measure in one dimension +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We have a general theory of covering theorems for doubling measures, developed notably in `density_theorems.lean`. In this file, we expand the API for this theory in one dimension, by showing that intervals belong to the relevant Vitali family. diff --git a/src/measure_theory/decomposition/lebesgue.lean b/src/measure_theory/decomposition/lebesgue.lean index eb1a19d20703b..b6c8c243b82d2 100644 --- a/src/measure_theory/decomposition/lebesgue.lean +++ b/src/measure_theory/decomposition/lebesgue.lean @@ -12,6 +12,9 @@ import measure_theory.function.ae_eq_of_integral /-! # Lebesgue decomposition +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves the Lebesgue decomposition theorem. The Lebesgue decomposition theorem states that, given two σ-finite measures `μ` and `ν`, there exists a σ-finite measure `ξ` and a measurable function `f` such that `μ = ξ + fν` and `ξ` is mutually singular with respect to `ν`. diff --git a/src/measure_theory/decomposition/radon_nikodym.lean b/src/measure_theory/decomposition/radon_nikodym.lean index ce7f41201fe0e..a6cfab57e8e65 100644 --- a/src/measure_theory/decomposition/radon_nikodym.lean +++ b/src/measure_theory/decomposition/radon_nikodym.lean @@ -8,6 +8,9 @@ import measure_theory.decomposition.lebesgue /-! # Radon-Nikodym theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves the Radon-Nikodym theorem. The Radon-Nikodym theorem states that, given measures `μ, ν`, if `have_lebesgue_decomposition μ ν`, then `μ` is absolutely continuous with respect to `ν` if and only if there exists a measurable function `f : α → ℝ≥0∞` such that `μ = fν`. diff --git a/src/measure_theory/function/continuous_map_dense.lean b/src/measure_theory/function/continuous_map_dense.lean index ab07dc7c853a4..c7c511cfef3a1 100644 --- a/src/measure_theory/function/continuous_map_dense.lean +++ b/src/measure_theory/function/continuous_map_dense.lean @@ -12,6 +12,9 @@ import measure_theory.integral.bochner /-! # Approximation in Lᵖ by continuous functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves that bounded continuous functions are dense in `Lp E p μ`, for `p < ∞`, if the domain `α` of the functions is a normal topological space and the measure `μ` is weakly regular. It also proves the same results for approximation by continuous functions with compact support diff --git a/src/measure_theory/function/l2_space.lean b/src/measure_theory/function/l2_space.lean index 6ab58ef0c4a86..73e7a16a5e4d8 100644 --- a/src/measure_theory/function/l2_space.lean +++ b/src/measure_theory/function/l2_space.lean @@ -9,6 +9,9 @@ import measure_theory.integral.set_integral /-! # `L^2` space +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + If `E` is an inner product space over `𝕜` (`ℝ` or `ℂ`), then `Lp E 2 μ` (defined in `lp_space.lean`) is also an inner product space, with inner product defined as `inner f g = ∫ a, ⟪f a, g a⟫ ∂μ`. diff --git a/src/measure_theory/function/special_functions/arctan.lean b/src/measure_theory/function/special_functions/arctan.lean index d1fdd51f129b3..2d10fb0a4c189 100644 --- a/src/measure_theory/function/special_functions/arctan.lean +++ b/src/measure_theory/function/special_functions/arctan.lean @@ -10,6 +10,9 @@ import measure_theory.constructions.borel_space.basic /-! # Measurability of arctan +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + -/ namespace real diff --git a/src/measure_theory/integral/interval_average.lean b/src/measure_theory/integral/interval_average.lean index 7b7f651c4ddd7..14e8f186367c2 100644 --- a/src/measure_theory/integral/interval_average.lean +++ b/src/measure_theory/integral/interval_average.lean @@ -9,6 +9,9 @@ import measure_theory.integral.average /-! # Integral average over an interval +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we introduce notation `⨍ x in a..b, f x` for the average `⨍ x in Ι a b, f x` of `f` over the interval `Ι a b = set.Ioc (min a b) (max a b)` w.r.t. the Lebesgue measure, then prove formulas for this average: diff --git a/src/measure_theory/integral/interval_integral.lean b/src/measure_theory/integral/interval_integral.lean index 9c446ff60f04a..e33e59b0144fc 100644 --- a/src/measure_theory/integral/interval_integral.lean +++ b/src/measure_theory/integral/interval_integral.lean @@ -10,6 +10,9 @@ import measure_theory.measure.lebesgue.basic /-! # Integral over an interval +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define `∫ x in a..b, f x ∂μ` to be `∫ x in Ioc a b, f x ∂μ` if `a ≤ b` and `-∫ x in Ioc b a, f x ∂μ` if `b ≤ a`. diff --git a/src/measure_theory/integral/peak_function.lean b/src/measure_theory/integral/peak_function.lean index 27c8bb7266666..0436fb4cce815 100644 --- a/src/measure_theory/integral/peak_function.lean +++ b/src/measure_theory/integral/peak_function.lean @@ -9,6 +9,9 @@ import measure_theory.function.locally_integrable /-! # Integrals against peak functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A sequence of peak functions is a sequence of functions with average one concentrating around a point `x₀`. Given such a sequence `φₙ`, then `∫ φₙ g` tends to `g x₀` in many situations, with a whole zoo of possible assumptions on `φₙ` and `g`. This file is devoted to such results. diff --git a/src/measure_theory/measure/finite_measure.lean b/src/measure_theory/measure/finite_measure.lean index 2f56da55396ca..3efd56a6070fe 100644 --- a/src/measure_theory/measure/finite_measure.lean +++ b/src/measure_theory/measure/finite_measure.lean @@ -10,6 +10,9 @@ import measure_theory.integral.bochner /-! # Finite measures +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the type of finite measures on a given measurable space. When the underlying space has a topology and the measurable space structure (sigma algebra) is finer than the Borel sigma algebra, then the type of finite measures is equipped with the topology of weak convergence diff --git a/src/measure_theory/measure/haar/quotient.lean b/src/measure_theory/measure/haar/quotient.lean index 87116a3ea240e..03f20400f0398 100644 --- a/src/measure_theory/measure/haar/quotient.lean +++ b/src/measure_theory/measure/haar/quotient.lean @@ -11,6 +11,9 @@ import algebra.group.opposite /-! # Haar quotient measure +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we consider properties of fundamental domains and measures for the action of a subgroup of a group `G` on `G` itself. diff --git a/src/measure_theory/measure/probability_measure.lean b/src/measure_theory/measure/probability_measure.lean index 8e79c08d4f447..f7222221eacd9 100644 --- a/src/measure_theory/measure/probability_measure.lean +++ b/src/measure_theory/measure/probability_measure.lean @@ -9,6 +9,9 @@ import measure_theory.integral.average /-! # Probability measures +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the type of probability measures on a given measurable space. When the underlying space has a topology and the measurable space structure (sigma algebra) is finer than the Borel sigma algebra, then the type of probability measures is equipped with the topology of convergence diff --git a/src/probability/integration.lean b/src/probability/integration.lean index 78a17d30219de..104779460cd3e 100644 --- a/src/probability/integration.lean +++ b/src/probability/integration.lean @@ -9,6 +9,9 @@ import probability.independence.basic /-! # Integration in Probability Theory +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Integration results for independent random variables. Specifically, for two independent random variables X and Y over the extended non-negative reals, `E[X * Y] = E[X] * E[Y]`, and similar results. diff --git a/src/ring_theory/class_group.lean b/src/ring_theory/class_group.lean index 446a8b7cdbbb8..e6717c1bbf8bd 100644 --- a/src/ring_theory/class_group.lean +++ b/src/ring_theory/class_group.lean @@ -10,6 +10,9 @@ import ring_theory.dedekind_domain.ideal /-! # The ideal class group +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the ideal class group `class_group R` of fractional ideals of `R` inside its field of fractions. diff --git a/src/ring_theory/roots_of_unity/basic.lean b/src/ring_theory/roots_of_unity/basic.lean index 3c0bd5d6c363e..1af2919d9c743 100644 --- a/src/ring_theory/roots_of_unity/basic.lean +++ b/src/ring_theory/roots_of_unity/basic.lean @@ -18,6 +18,9 @@ import tactic.zify /-! # Roots of unity and primitive roots of unity +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define roots of unity in the context of an arbitrary commutative monoid, as a subgroup of the group of units. We also define a predicate `is_primitive_root` on commutative monoids, expressing that an element is a primitive root of unity. diff --git a/src/ring_theory/roots_of_unity/complex.lean b/src/ring_theory/roots_of_unity/complex.lean index af0185b342e9c..d5841c9d4726b 100644 --- a/src/ring_theory/roots_of_unity/complex.lean +++ b/src/ring_theory/roots_of_unity/complex.lean @@ -9,6 +9,9 @@ import ring_theory.roots_of_unity.basic /-! # Complex roots of unity +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we show that the `n`-th complex roots of unity are exactly the complex numbers `e ^ (2 * real.pi * complex.I * (i / n))` for `i ∈ finset.range n`. diff --git a/src/topology/order/hom/esakia.lean b/src/topology/order/hom/esakia.lean index ad18a06dbb155..1ae1c28310069 100644 --- a/src/topology/order/hom/esakia.lean +++ b/src/topology/order/hom/esakia.lean @@ -9,6 +9,9 @@ import topology.order.hom.basic /-! # Esakia morphisms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines pseudo-epimorphisms and Esakia morphisms. We use the `fun_like` design, so each type of morphisms has a companion typeclass which is meant to From 74f1d61944a5a793e8c939d47608178c0a0cb0c2 Mon Sep 17 00:00:00 2001 From: Apurva Date: Fri, 9 Jun 2023 01:23:28 +0000 Subject: [PATCH 1159/1291] feat(analysis/convex/cone/proper): define proper cone (#18913) Part of #16266 We define proper cones as nonempty, closed cones and define duals. Next todo: Prove Farkas' lemma from #19008 Co-authored-by: Apurva Nakade --- src/analysis/convex/cone/proper.lean | 143 ++++++++++++++++++++++++--- 1 file changed, 130 insertions(+), 13 deletions(-) diff --git a/src/analysis/convex/cone/proper.lean b/src/analysis/convex/cone/proper.lean index e210123b9d5e6..e85a633104328 100644 --- a/src/analysis/convex/cone/proper.lean +++ b/src/analysis/convex/cone/proper.lean @@ -3,10 +3,10 @@ Copyright (c) 2022 Apurva Nakade All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Apurva Nakade -/ -import analysis.convex.cone.basic -import topology.algebra.monoid +import analysis.convex.cone.dual /-! +# Proper cones We define a proper cone as a nonempty, closed, convex cone. Proper cones are used in defining conic programs which generalize linear programs. A linear program is a conic program for the positive @@ -16,13 +16,16 @@ linear programs, the results from this file can be used to prove duality theorem ## TODO -In the next few PRs (already sorry-free), we will add the definition and prove several properties -of proper cones and finally prove the cone version of Farkas' lemma (2.3.4 in the reference). - The next steps are: +- Prove the cone version of Farkas' lemma (2.3.4 in the reference). +- Add comap, adjoint +- Add convex_cone_class that extends set_like and replace the below instance +- Define the positive cone as a proper cone. - Define primal and dual cone programs and prove weak duality. - Prove regular and strong duality for cone programs using Farkas' lemma (see reference). - Define linear programs and prove LP duality as a special case of cone duality. +- Find a better reference (textbook instead of lecture notes). +- Show submodules are (proper) cones. ## References @@ -32,21 +35,135 @@ The next steps are: namespace convex_cone -variables {E : Type*} [add_comm_monoid E] [has_smul ℝ E] [topological_space E] - [has_continuous_const_smul ℝ E] [has_continuous_add E] +variables {𝕜 : Type*} [ordered_semiring 𝕜] +variables {E : Type*} [add_comm_monoid E] [topological_space E] [has_continuous_add E] + [has_smul 𝕜 E] [has_continuous_const_smul 𝕜 E] -/-- The closure of a convex cone inside a real inner product space is a convex cone. This +/-- The closure of a convex cone inside a topological space as a convex cone. This construction is mainly used for defining maps between proper cones. -/ -protected def closure (K : convex_cone ℝ E) : convex_cone ℝ E := +protected def closure (K : convex_cone 𝕜 E) : convex_cone 𝕜 E := { carrier := closure ↑K, smul_mem' := λ c hc _ h₁, map_mem_closure (continuous_id'.const_smul c) h₁ (λ _ h₂, K.smul_mem hc h₂), add_mem' := λ _ h₁ _ h₂, map_mem_closure₂ continuous_add h₁ h₂ K.add_mem } -@[simp, norm_cast] lemma coe_closure (K : convex_cone ℝ E) : (K.closure : set E) = closure K := rfl +@[simp, norm_cast] lemma coe_closure (K : convex_cone 𝕜 E) : (K.closure : set E) = closure K := rfl + +@[simp] protected lemma mem_closure {K : convex_cone 𝕜 E} {a : E} : + a ∈ K.closure ↔ a ∈ closure (K : set E) := iff.rfl -protected lemma mem_closure {K : convex_cone ℝ E} {a : E} : - a ∈ K.closure ↔ a ∈ closure (K : set E) := -iff.rfl +@[simp] lemma closure_eq {K L : convex_cone 𝕜 E} : K.closure = L ↔ closure (K : set E) = L := +set_like.ext'_iff end convex_cone + +/-- A proper cone is a convex cone `K` that is nonempty and closed. Proper cones have the nice +property that the dual of the dual of a proper cone is itself. This makes them useful for defining +cone programs and proving duality theorems. -/ +structure proper_cone (𝕜 : Type*) (E : Type*) + [ordered_semiring 𝕜] [add_comm_monoid E] [topological_space E] [has_smul 𝕜 E] + extends convex_cone 𝕜 E := +(nonempty' : (carrier : set E).nonempty) +(is_closed' : is_closed (carrier : set E)) + +namespace proper_cone + +section has_smul + +variables {𝕜 : Type*} [ordered_semiring 𝕜] +variables {E : Type*} [add_comm_monoid E] [topological_space E] [has_smul 𝕜 E] + +instance : has_coe (proper_cone 𝕜 E) (convex_cone 𝕜 E) := ⟨λ K, K.1⟩ + +@[simp] lemma to_convex_cone_eq_coe (K : proper_cone 𝕜 E) : K.to_convex_cone = K := rfl + +lemma ext' : function.injective (coe : proper_cone 𝕜 E → convex_cone 𝕜 E) := +λ S T h, by cases S; cases T; congr' + +-- TODO: add convex_cone_class that extends set_like and replace the below instance +instance : set_like (proper_cone 𝕜 E) E := +{ coe := λ K, K.carrier, + coe_injective' := λ _ _ h, proper_cone.ext' (set_like.coe_injective h) } + +@[ext] lemma ext {S T : proper_cone 𝕜 E} (h : ∀ x, x ∈ S ↔ x ∈ T) : S = T := set_like.ext h + +@[simp] lemma mem_coe {x : E} {K : proper_cone 𝕜 E} : x ∈ (K : convex_cone 𝕜 E) ↔ x ∈ K := iff.rfl + +protected lemma nonempty (K : proper_cone 𝕜 E) : (K : set E).nonempty := K.nonempty' + +protected lemma is_closed (K : proper_cone 𝕜 E) : is_closed (K : set E) := K.is_closed' + +end has_smul + +section module + +variables {𝕜 : Type*} [ordered_semiring 𝕜] +variables {E : Type*} [add_comm_monoid E] [topological_space E] [t1_space E] [module 𝕜 E] + +instance : has_zero (proper_cone 𝕜 E) := +⟨ { to_convex_cone := 0, + nonempty' := ⟨0, rfl⟩, + is_closed' := is_closed_singleton } ⟩ + +instance : inhabited (proper_cone 𝕜 E) := ⟨0⟩ + +@[simp] lemma mem_zero (x : E) : x ∈ (0 : proper_cone 𝕜 E) ↔ x = 0 := iff.rfl +@[simp, norm_cast] lemma coe_zero : ↑(0 : proper_cone 𝕜 E) = (0 : convex_cone 𝕜 E) := rfl + +lemma pointed_zero : (0 : proper_cone 𝕜 E).pointed := by simp [convex_cone.pointed_zero] + +end module + +section inner_product_space + +variables {E : Type*} [normed_add_comm_group E] [inner_product_space ℝ E] +variables {F : Type*} [normed_add_comm_group F] [inner_product_space ℝ F] + +protected lemma pointed (K : proper_cone ℝ E) : (K : convex_cone ℝ E).pointed := +(K : convex_cone ℝ E).pointed_of_nonempty_of_is_closed K.nonempty K.is_closed + +/-- The closure of image of a proper cone under a continuous `ℝ`-linear map is a proper cone. We +use continuous maps here so that the adjoint of f is also a map between proper cones. -/ +noncomputable def map (f : E →L[ℝ] F) (K : proper_cone ℝ E) : proper_cone ℝ F := +{ to_convex_cone := convex_cone.closure (convex_cone.map (f : E →ₗ[ℝ] F) ↑K), + nonempty' := ⟨ 0, subset_closure $ set_like.mem_coe.2 $ convex_cone.mem_map.2 + ⟨0, K.pointed, map_zero _⟩ ⟩, + is_closed' := is_closed_closure } + +@[simp, norm_cast] lemma coe_map (f : E →L[ℝ] F) (K : proper_cone ℝ E) : + ↑(K.map f) = (convex_cone.map (f : E →ₗ[ℝ] F) ↑K).closure := rfl + +@[simp] lemma mem_map {f : E →L[ℝ] F} {K : proper_cone ℝ E} {y : F} : + y ∈ K.map f ↔ y ∈ (convex_cone.map (f : E →ₗ[ℝ] F) ↑K).closure := iff.rfl + +@[simp] lemma map_id (K : proper_cone ℝ E) : K.map (continuous_linear_map.id ℝ E) = K := +proper_cone.ext' $ by simpa using is_closed.closure_eq K.is_closed + +/-- The inner dual cone of a proper cone is a proper cone. -/ +def dual (K : proper_cone ℝ E): (proper_cone ℝ E) := +{ to_convex_cone := (K : set E).inner_dual_cone, + nonempty' := ⟨0, pointed_inner_dual_cone _⟩, + is_closed' := is_closed_inner_dual_cone _ } + +@[simp, norm_cast] +lemma coe_dual (K : proper_cone ℝ E) : ↑(dual K) = (K : set E).inner_dual_cone := rfl + +@[simp] lemma mem_dual {K : proper_cone ℝ E} {y : E} : + y ∈ dual K ↔ ∀ ⦃x⦄, x ∈ K → 0 ≤ ⟪x, y⟫_ℝ := +by {rw [← mem_coe, coe_dual, mem_inner_dual_cone _ _], refl} + +-- TODO: add comap, adjoint + +end inner_product_space + +section complete_space + +variables {E : Type*} [normed_add_comm_group E] [inner_product_space ℝ E] [complete_space E] + +/-- The dual of the dual of a proper cone is itself. -/ +@[simp] theorem dual_dual (K : proper_cone ℝ E) : K.dual.dual = K := proper_cone.ext' $ + (K : convex_cone ℝ E).inner_dual_cone_of_inner_dual_cone_eq_self K.nonempty K.is_closed + +end complete_space + +end proper_cone From 6b31d1eebd64eab86d5bd9936bfaada6ca8b5842 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Fri, 9 Jun 2023 05:37:49 +0000 Subject: [PATCH 1160/1291] chore(*): add mathlib4 synchronization comments (#19167) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.algebra.spectrum` * `algebra.category.Module.algebra` * `algebra.lie.free` * `algebra.lie.normalizer` * `algebra.lie.quotient` * `algebra.lie.universal_enveloping` * `algebra.order.complete_field` * `algebra.order.interval` * `algebra.star.order` * `analysis.normed_space.enorm` * `analysis.special_functions.integrals` * `analysis.special_functions.non_integrable` * `analysis.special_functions.polar_coord` * `category_theory.monoidal.of_has_finite_products` * `category_theory.monoidal.opposite` * `category_theory.monoidal.skeleton` * `category_theory.monoidal.types.coyoneda` * `category_theory.preadditive.Mat` * `data.rat.star` * `data.real.pi.wallis` * `field_theory.adjoin` * `group_theory.schur_zassenhaus` * `linear_algebra.eigenspace.basic` * `linear_algebra.tensor_algebra.grading` * `measure_theory.covering.besicovitch` * `measure_theory.covering.besicovitch_vector_space` * `measure_theory.function.jacobian` * `measure_theory.integral.divergence_theorem` * `measure_theory.integral.fund_thm_calculus` * `measure_theory.measure.lebesgue.integral` * `number_theory.ramification_inertia` * `ring_theory.dedekind_domain.dvr` * `ring_theory.dedekind_domain.pid` * `ring_theory.discrete_valuation_ring.tfae` * `ring_theory.polynomial.cyclotomic.basic` * `ring_theory.witt_vector.basic` * `ring_theory.witt_vector.is_poly` * `ring_theory.witt_vector.mul_p` * `ring_theory.witt_vector.teichmuller` * `topology.algebra.continuous_affine_map` * `topology.algebra.module.character_space` * `topology.category.Compactum` * `topology.continuous_function.units` --- src/algebra/algebra/spectrum.lean | 3 +++ src/algebra/category/Module/algebra.lean | 3 +++ src/algebra/lie/free.lean | 3 +++ src/algebra/lie/normalizer.lean | 3 +++ src/algebra/lie/quotient.lean | 3 +++ src/algebra/lie/universal_enveloping.lean | 3 +++ src/algebra/order/complete_field.lean | 3 +++ src/algebra/order/interval.lean | 3 +++ src/algebra/star/order.lean | 3 +++ src/analysis/normed_space/enorm.lean | 3 +++ src/analysis/special_functions/integrals.lean | 3 +++ src/analysis/special_functions/non_integrable.lean | 3 +++ src/analysis/special_functions/polar_coord.lean | 3 +++ src/category_theory/monoidal/of_has_finite_products.lean | 3 +++ src/category_theory/monoidal/opposite.lean | 3 +++ src/category_theory/monoidal/skeleton.lean | 3 +++ src/category_theory/monoidal/types/coyoneda.lean | 3 +++ src/category_theory/preadditive/Mat.lean | 3 +++ src/data/rat/star.lean | 3 +++ src/data/real/pi/wallis.lean | 3 +++ src/field_theory/adjoin.lean | 3 +++ src/group_theory/schur_zassenhaus.lean | 3 +++ src/linear_algebra/eigenspace/basic.lean | 3 +++ src/linear_algebra/tensor_algebra/grading.lean | 3 +++ src/measure_theory/covering/besicovitch.lean | 3 +++ src/measure_theory/covering/besicovitch_vector_space.lean | 3 +++ src/measure_theory/function/jacobian.lean | 3 +++ src/measure_theory/integral/divergence_theorem.lean | 3 +++ src/measure_theory/integral/fund_thm_calculus.lean | 3 +++ src/measure_theory/measure/lebesgue/integral.lean | 5 ++++- src/number_theory/ramification_inertia.lean | 3 +++ src/ring_theory/dedekind_domain/dvr.lean | 3 +++ src/ring_theory/dedekind_domain/pid.lean | 3 +++ src/ring_theory/discrete_valuation_ring/tfae.lean | 3 +++ src/ring_theory/polynomial/cyclotomic/basic.lean | 3 +++ src/ring_theory/witt_vector/basic.lean | 3 +++ src/ring_theory/witt_vector/is_poly.lean | 3 +++ src/ring_theory/witt_vector/mul_p.lean | 3 +++ src/ring_theory/witt_vector/teichmuller.lean | 3 +++ src/topology/algebra/continuous_affine_map.lean | 3 +++ src/topology/algebra/module/character_space.lean | 3 +++ src/topology/category/Compactum.lean | 3 +++ src/topology/continuous_function/units.lean | 3 +++ 43 files changed, 130 insertions(+), 1 deletion(-) diff --git a/src/algebra/algebra/spectrum.lean b/src/algebra/algebra/spectrum.lean index cdbf09f8e3bd4..6b5043020159e 100644 --- a/src/algebra/algebra/spectrum.lean +++ b/src/algebra/algebra/spectrum.lean @@ -8,6 +8,9 @@ import algebra.star.subalgebra import tactic.noncomm_ring /-! # Spectrum of an element in an algebra + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. This file develops the basic theory of the spectrum of an element of an algebra. This theory will serve as the foundation for spectral theory in Banach algebras. diff --git a/src/algebra/category/Module/algebra.lean b/src/algebra/category/Module/algebra.lean index 224baa250b672..ba74d2d24bfb0 100644 --- a/src/algebra/category/Module/algebra.lean +++ b/src/algebra/category/Module/algebra.lean @@ -10,6 +10,9 @@ import algebra.category.Module.basic /-! # Additional typeclass for modules over an algebra +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + For an object in `M : Module A`, where `A` is a `k`-algebra, we provide additional typeclasses on the underlying type `M`, namely `module k M` and `is_scalar_tower k A M`. diff --git a/src/algebra/lie/free.lean b/src/algebra/lie/free.lean index 84671c2e4e28a..1823531a53c2a 100644 --- a/src/algebra/lie/free.lean +++ b/src/algebra/lie/free.lean @@ -11,6 +11,9 @@ import algebra.free_non_unital_non_assoc_algebra /-! # Free Lie algebras +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given a commutative ring `R` and a type `X` we construct the free Lie algebra on `X` with coefficients in `R` together with its universal property. diff --git a/src/algebra/lie/normalizer.lean b/src/algebra/lie/normalizer.lean index e5a30f89fef4d..b8b889787bf49 100644 --- a/src/algebra/lie/normalizer.lean +++ b/src/algebra/lie/normalizer.lean @@ -10,6 +10,9 @@ import algebra.lie.quotient /-! # The normalizer of a Lie submodules and subalgebras. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given a Lie module `M` over a Lie subalgebra `L`, the normalizer of a Lie submodule `N ⊆ M` is the Lie submodule with underlying set `{ m | ∀ (x : L), ⁅x, m⁆ ∈ N }`. diff --git a/src/algebra/lie/quotient.lean b/src/algebra/lie/quotient.lean index ad05865053818..16cb98882faaf 100644 --- a/src/algebra/lie/quotient.lean +++ b/src/algebra/lie/quotient.lean @@ -10,6 +10,9 @@ import linear_algebra.isomorphisms /-! # Quotients of Lie algebras and Lie modules +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given a Lie submodule of a Lie module, the quotient carries a natural Lie module structure. In the special case that the Lie module is the Lie algebra itself via the adjoint action, the submodule is a Lie ideal and the quotient carries a natural Lie algebra structure. diff --git a/src/algebra/lie/universal_enveloping.lean b/src/algebra/lie/universal_enveloping.lean index 218f402cb80cd..b9a3dbd8834e0 100644 --- a/src/algebra/lie/universal_enveloping.lean +++ b/src/algebra/lie/universal_enveloping.lean @@ -10,6 +10,9 @@ import linear_algebra.tensor_algebra.basic /-! # Universal enveloping algebra +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given a commutative ring `R` and a Lie algebra `L` over `R`, we construct the universal enveloping algebra of `L`, together with its universal property. diff --git a/src/algebra/order/complete_field.lean b/src/algebra/order/complete_field.lean index 5e62fce7c5440..d4d207ca1a804 100644 --- a/src/algebra/order/complete_field.lean +++ b/src/algebra/order/complete_field.lean @@ -10,6 +10,9 @@ import analysis.special_functions.pow.real /-! # Conditionally complete linear ordered fields +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file shows that the reals are unique, or, more formally, given a type satisfying the common axioms of the reals (field, conditionally complete, linearly ordered) that there is an isomorphism preserving these properties to the reals. This is `rat.induced_order_ring_iso`. Moreover this diff --git a/src/algebra/order/interval.lean b/src/algebra/order/interval.lean index 77cdf30d29b8e..5372b31fabdea 100644 --- a/src/algebra/order/interval.lean +++ b/src/algebra/order/interval.lean @@ -13,6 +13,9 @@ import tactic.positivity /-! # Interval arithmetic +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines arithmetic operations on intervals and prove their correctness. Note that this is full precision operations. The essentials of float operations can be found in `data.fp.basic`. We have not yet integrated these with the rest of the library. diff --git a/src/algebra/star/order.lean b/src/algebra/star/order.lean index 540a7acd6351f..7cee569d67e6f 100644 --- a/src/algebra/star/order.lean +++ b/src/algebra/star/order.lean @@ -9,6 +9,9 @@ import group_theory.submonoid.basic /-! # Star ordered rings +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define the class `star_ordered_ring R`, which says that the order on `R` respects the star operation, i.e. an element `r` is nonnegative iff it is in the `add_submonoid` generated by elements of the form `star s * s`. In many cases, including all C⋆-algebras, this can be reduced to diff --git a/src/analysis/normed_space/enorm.lean b/src/analysis/normed_space/enorm.lean index b355dfce88d26..c5aeb25999aff 100644 --- a/src/analysis/normed_space/enorm.lean +++ b/src/analysis/normed_space/enorm.lean @@ -8,6 +8,9 @@ import analysis.normed_space.basic /-! # Extended norm +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define a structure `enorm 𝕜 V` representing an extended norm (i.e., a norm that can take the value `∞`) on a vector space `V` over a normed field `𝕜`. We do not use `class` for an `enorm` because the same space can have more than one extended norm. For example, the space of diff --git a/src/analysis/special_functions/integrals.lean b/src/analysis/special_functions/integrals.lean index 9291d26b5273a..22b8ec9e9bff1 100644 --- a/src/analysis/special_functions/integrals.lean +++ b/src/analysis/special_functions/integrals.lean @@ -9,6 +9,9 @@ import analysis.special_functions.trigonometric.arctan_deriv /-! # Integration of specific interval integrals +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains proofs of the integrals of various specific functions. This includes: * Integrals of simple functions, such as `id`, `pow`, `inv`, `exp`, `log` * Integrals of some trigonometric functions, such as `sin`, `cos`, `1 / (1 + x^2)` diff --git a/src/analysis/special_functions/non_integrable.lean b/src/analysis/special_functions/non_integrable.lean index 387f073c9f762..d193676c5f1ae 100644 --- a/src/analysis/special_functions/non_integrable.lean +++ b/src/analysis/special_functions/non_integrable.lean @@ -9,6 +9,9 @@ import measure_theory.integral.fund_thm_calculus /-! # Non integrable functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove that the derivative of a function that tends to infinity is not interval integrable, see `interval_integral.not_integrable_has_deriv_at_of_tendsto_norm_at_top_filter` and `interval_integral.not_integrable_has_deriv_at_of_tendsto_norm_at_top_punctured`. Then we apply the diff --git a/src/analysis/special_functions/polar_coord.lean b/src/analysis/special_functions/polar_coord.lean index 18dbbde7ff807..93ff2a25343ea 100644 --- a/src/analysis/special_functions/polar_coord.lean +++ b/src/analysis/special_functions/polar_coord.lean @@ -9,6 +9,9 @@ import measure_theory.function.jacobian /-! # Polar coordinates +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define polar coordinates, as a local homeomorphism in `ℝ^2` between `ℝ^2 - (-∞, 0]` and `(0, +∞) × (-π, π)`. Its inverse is given by `(r, θ) ↦ (r cos θ, r sin θ)`. diff --git a/src/category_theory/monoidal/of_has_finite_products.lean b/src/category_theory/monoidal/of_has_finite_products.lean index 7ec14943a15dd..593daf66b24e1 100644 --- a/src/category_theory/monoidal/of_has_finite_products.lean +++ b/src/category_theory/monoidal/of_has_finite_products.lean @@ -10,6 +10,9 @@ import category_theory.limits.shapes.terminal /-! # The natural monoidal structure on any category with finite (co)products. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A category with a monoidal structure provided in this way is sometimes called a (co)cartesian category, although this is also sometimes used to mean a finitely complete category. diff --git a/src/category_theory/monoidal/opposite.lean b/src/category_theory/monoidal/opposite.lean index fcb952d416984..ff8730fff478e 100644 --- a/src/category_theory/monoidal/opposite.lean +++ b/src/category_theory/monoidal/opposite.lean @@ -8,6 +8,9 @@ import category_theory.monoidal.coherence /-! # Monoidal opposites +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We write `Cᵐᵒᵖ` for the monoidal opposite of a monoidal category `C`. -/ diff --git a/src/category_theory/monoidal/skeleton.lean b/src/category_theory/monoidal/skeleton.lean index f842f1efcedc6..ab6b9cdf4fd4b 100644 --- a/src/category_theory/monoidal/skeleton.lean +++ b/src/category_theory/monoidal/skeleton.lean @@ -10,6 +10,9 @@ import category_theory.skeletal /-! # The monoid on the skeleton of a monoidal category +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The skeleton of a monoidal category is a monoid. -/ diff --git a/src/category_theory/monoidal/types/coyoneda.lean b/src/category_theory/monoidal/types/coyoneda.lean index f9204ee86ec1c..9d1b1fa6ae7db 100644 --- a/src/category_theory/monoidal/types/coyoneda.lean +++ b/src/category_theory/monoidal/types/coyoneda.lean @@ -8,6 +8,9 @@ import category_theory.monoidal.coherence_lemmas /-! # `(𝟙_ C ⟶ -)` is a lax monoidal functor to `Type` + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ open category_theory diff --git a/src/category_theory/preadditive/Mat.lean b/src/category_theory/preadditive/Mat.lean index 96a248a0caa20..3794457cb5e65 100644 --- a/src/category_theory/preadditive/Mat.lean +++ b/src/category_theory/preadditive/Mat.lean @@ -17,6 +17,9 @@ import algebra.opposites /-! # Matrices over a category. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + When `C` is a preadditive category, `Mat_ C` is the preadditive category whose objects are finite tuples of objects in `C`, and whose morphisms are matrices of morphisms from `C`. diff --git a/src/data/rat/star.lean b/src/data/rat/star.lean index e254482fe566e..3d999719221c8 100644 --- a/src/data/rat/star.lean +++ b/src/data/rat/star.lean @@ -10,6 +10,9 @@ import group_theory.submonoid.membership /-! # Star order structure on ℚ +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Here we put the trivial `star` operation on `ℚ` for convenience and show that it is a `star_ordered_ring`. In particular, this means that every element of `ℚ` is a sum of squares. -/ diff --git a/src/data/real/pi/wallis.lean b/src/data/real/pi/wallis.lean index 2f114d5606c4e..8869286478aa2 100644 --- a/src/data/real/pi/wallis.lean +++ b/src/data/real/pi/wallis.lean @@ -7,6 +7,9 @@ import analysis.special_functions.integrals /-! # The Wallis formula for Pi +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file establishes the Wallis product for `π` (`real.tendsto_prod_pi_div_two`). Our proof is largely about analyzing the behaviour of the sequence `∫ x in 0..π, sin x ^ n` as `n → ∞`. See: https://en.wikipedia.org/wiki/Wallis_product diff --git a/src/field_theory/adjoin.lean b/src/field_theory/adjoin.lean index 13e44e04c536d..24cc89792d68f 100644 --- a/src/field_theory/adjoin.lean +++ b/src/field_theory/adjoin.lean @@ -12,6 +12,9 @@ import ring_theory.tensor_product /-! # Adjoining Elements to Fields +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we introduce the notion of adjoining elements to fields. This isn't quite the same as adjoining elements to rings. For example, `algebra.adjoin K {x}` might not include `x⁻¹`. diff --git a/src/group_theory/schur_zassenhaus.lean b/src/group_theory/schur_zassenhaus.lean index 58c68bd95b0b2..fffa1ed9e956b 100644 --- a/src/group_theory/schur_zassenhaus.lean +++ b/src/group_theory/schur_zassenhaus.lean @@ -10,6 +10,9 @@ import group_theory.transfer /-! # The Schur-Zassenhaus Theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove the Schur-Zassenhaus theorem. ## Main results diff --git a/src/linear_algebra/eigenspace/basic.lean b/src/linear_algebra/eigenspace/basic.lean index fcfe2caf3a4bf..8de6594e40b3b 100644 --- a/src/linear_algebra/eigenspace/basic.lean +++ b/src/linear_algebra/eigenspace/basic.lean @@ -12,6 +12,9 @@ import linear_algebra.finite_dimensional /-! # Eigenvectors and eigenvalues +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines eigenspaces, eigenvalues, and eigenvalues, as well as their generalized counterparts. We follow Axler's approach [axler2015] because it allows us to derive many properties without choosing a basis and without using matrices. diff --git a/src/linear_algebra/tensor_algebra/grading.lean b/src/linear_algebra/tensor_algebra/grading.lean index e51bdd02f6441..9452c252cb44c 100644 --- a/src/linear_algebra/tensor_algebra/grading.lean +++ b/src/linear_algebra/tensor_algebra/grading.lean @@ -9,6 +9,9 @@ import ring_theory.graded_algebra.basic /-! # Results about the grading structure of the tensor algebra +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The main result is `tensor_algebra.graded_algebra`, which says that the tensor algebra is a ℕ-graded algebra. -/ diff --git a/src/measure_theory/covering/besicovitch.lean b/src/measure_theory/covering/besicovitch.lean index 9776ab4c2b7f3..55bc5ce327c6e 100644 --- a/src/measure_theory/covering/besicovitch.lean +++ b/src/measure_theory/covering/besicovitch.lean @@ -13,6 +13,9 @@ import topology.metric_space.basic /-! # Besicovitch covering theorems +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The topological Besicovitch covering theorem ensures that, in a nice metric space, there exists a number `N` such that, from any family of balls with bounded radii, one can extract `N` families, each made of disjoint balls, covering together all the centers of the initial family. diff --git a/src/measure_theory/covering/besicovitch_vector_space.lean b/src/measure_theory/covering/besicovitch_vector_space.lean index fab628357d0ab..130fe156abfde 100644 --- a/src/measure_theory/covering/besicovitch_vector_space.lean +++ b/src/measure_theory/covering/besicovitch_vector_space.lean @@ -10,6 +10,9 @@ import measure_theory.covering.besicovitch /-! # Satellite configurations for Besicovitch covering lemma in vector spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The Besicovitch covering theorem ensures that, in a nice metric space, there exists a number `N` such that, from any family of balls with bounded radii, one can extract `N` families, each made of disjoint balls, covering together all the centers of the initial family. diff --git a/src/measure_theory/function/jacobian.lean b/src/measure_theory/function/jacobian.lean index ce5efac87502b..c7469efe8f991 100644 --- a/src/measure_theory/function/jacobian.lean +++ b/src/measure_theory/function/jacobian.lean @@ -13,6 +13,9 @@ import measure_theory.constructions.polish /-! # Change of variables in higher-dimensional integrals +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Let `μ` be a Lebesgue measure on a finite-dimensional real vector space `E`. Let `f : E → E` be a function which is injective and differentiable on a measurable set `s`, with derivative `f'`. Then we prove that `f '' s` is measurable, and diff --git a/src/measure_theory/integral/divergence_theorem.lean b/src/measure_theory/integral/divergence_theorem.lean index 5ec6d78cadbce..5b52d7cf9ad36 100644 --- a/src/measure_theory/integral/divergence_theorem.lean +++ b/src/measure_theory/integral/divergence_theorem.lean @@ -12,6 +12,9 @@ import measure_theory.integral.interval_integral /-! # Divergence theorem for Bochner integral +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove the Divergence theorem for Bochner integral on a box in `ℝⁿ⁺¹ = fin (n + 1) → ℝ`. More precisely, we prove the following theorem. diff --git a/src/measure_theory/integral/fund_thm_calculus.lean b/src/measure_theory/integral/fund_thm_calculus.lean index d19811f6d32cb..5395634b24f97 100644 --- a/src/measure_theory/integral/fund_thm_calculus.lean +++ b/src/measure_theory/integral/fund_thm_calculus.lean @@ -15,6 +15,9 @@ import measure_theory.integral.vitali_caratheodory /-! # Fundamental Theorem of Calculus +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We prove various versions of the [fundamental theorem of calculus](https://en.wikipedia.org/wiki/Fundamental_theorem_of_calculus) for interval integrals in `ℝ`. diff --git a/src/measure_theory/measure/lebesgue/integral.lean b/src/measure_theory/measure/lebesgue/integral.lean index a7542af910607..126802eae5fd5 100644 --- a/src/measure_theory/measure/lebesgue/integral.lean +++ b/src/measure_theory/measure/lebesgue/integral.lean @@ -6,7 +6,10 @@ Authors: Johannes Hölzl, Sébastien Gouëzel, Yury Kudryashov import measure_theory.integral.set_integral import measure_theory.measure.lebesgue.basic -/-! # Properties of integration with respect to the Lebesgue measure -/ +/-! # Properties of integration with respect to the Lebesgue measure + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4.-/ open set filter measure_theory measure_theory.measure topological_space diff --git a/src/number_theory/ramification_inertia.lean b/src/number_theory/ramification_inertia.lean index 8c61439b6981f..91585ce4cc595 100644 --- a/src/number_theory/ramification_inertia.lean +++ b/src/number_theory/ramification_inertia.lean @@ -10,6 +10,9 @@ import ring_theory.dedekind_domain.ideal /-! # Ramification index and inertia degree +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given `P : ideal S` lying over `p : ideal R` for the ring extension `f : R →+* S` (assuming `P` and `p` are prime or maximal where needed), the **ramification index** `ideal.ramification_idx f p P` is the multiplicity of `P` in `map f p`, diff --git a/src/ring_theory/dedekind_domain/dvr.lean b/src/ring_theory/dedekind_domain/dvr.lean index 62b3fd3e1307b..4e4592bbdea22 100644 --- a/src/ring_theory/dedekind_domain/dvr.lean +++ b/src/ring_theory/dedekind_domain/dvr.lean @@ -10,6 +10,9 @@ import ring_theory.discrete_valuation_ring.tfae /-! # Dedekind domains +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines an equivalent notion of a Dedekind domain (or Dedekind ring), namely a Noetherian integral domain where the localization at all nonzero prime ideals is a DVR (TODO: and shows that implies the main definition). diff --git a/src/ring_theory/dedekind_domain/pid.lean b/src/ring_theory/dedekind_domain/pid.lean index fd41b7a7737fb..0cab71477b16b 100644 --- a/src/ring_theory/dedekind_domain/pid.lean +++ b/src/ring_theory/dedekind_domain/pid.lean @@ -10,6 +10,9 @@ import ring_theory.dedekind_domain.ideal /-! # Proving a Dedekind domain is a PID +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains some results that we can use to show all ideals in a Dedekind domain are principal. diff --git a/src/ring_theory/discrete_valuation_ring/tfae.lean b/src/ring_theory/discrete_valuation_ring/tfae.lean index bae86b2c29902..6beb45ac83a56 100644 --- a/src/ring_theory/discrete_valuation_ring/tfae.lean +++ b/src/ring_theory/discrete_valuation_ring/tfae.lean @@ -12,6 +12,9 @@ import ring_theory.nakayama # Equivalent conditions for DVR +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In `discrete_valuation_ring.tfae`, we show that the following are equivalent for a noetherian local domain `(R, m, k)`: - `R` is a discrete valuation ring diff --git a/src/ring_theory/polynomial/cyclotomic/basic.lean b/src/ring_theory/polynomial/cyclotomic/basic.lean index 53f1675d35b2f..31c0f26811d84 100644 --- a/src/ring_theory/polynomial/cyclotomic/basic.lean +++ b/src/ring_theory/polynomial/cyclotomic/basic.lean @@ -18,6 +18,9 @@ import ring_theory.roots_of_unity.basic /-! # Cyclotomic polynomials. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + For `n : ℕ` and an integral domain `R`, we define a modified version of the `n`-th cyclotomic polynomial with coefficients in `R`, denoted `cyclotomic' n R`, as `∏ (X - μ)`, where `μ` varies over the primitive `n`th roots of unity. If there is a primitive `n`th root of unity in `R` then diff --git a/src/ring_theory/witt_vector/basic.lean b/src/ring_theory/witt_vector/basic.lean index dcb9ba16ff989..bd5c1584a3451 100644 --- a/src/ring_theory/witt_vector/basic.lean +++ b/src/ring_theory/witt_vector/basic.lean @@ -11,6 +11,9 @@ import ring_theory.witt_vector.defs /-! # Witt vectors +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file verifies that the ring operations on `witt_vector p R` satisfy the axioms of a commutative ring. diff --git a/src/ring_theory/witt_vector/is_poly.lean b/src/ring_theory/witt_vector/is_poly.lean index d0dfef20a0fbd..5be39c921c856 100644 --- a/src/ring_theory/witt_vector/is_poly.lean +++ b/src/ring_theory/witt_vector/is_poly.lean @@ -11,6 +11,9 @@ import data.mv_polynomial.funext /-! # The `is_poly` predicate +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + `witt_vector.is_poly` is a (type-valued) predicate on functions `f : Π R, 𝕎 R → 𝕎 R`. It asserts that there is a family of polynomials `φ : ℕ → mv_polynomial ℕ ℤ`, such that the `n`th coefficient of `f x` is equal to `φ n` evaluated on the coefficients of `x`. diff --git a/src/ring_theory/witt_vector/mul_p.lean b/src/ring_theory/witt_vector/mul_p.lean index c8cc34f12bd59..750ab4c8d9e65 100644 --- a/src/ring_theory/witt_vector/mul_p.lean +++ b/src/ring_theory/witt_vector/mul_p.lean @@ -9,6 +9,9 @@ import ring_theory.witt_vector.is_poly /-! ## Multiplication by `n` in the ring of Witt vectors +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we show that multiplication by `n` in the ring of Witt vectors is a polynomial function. We then use this fact to show that the composition of Frobenius and Verschiebung is equal to multiplication by `p`. diff --git a/src/ring_theory/witt_vector/teichmuller.lean b/src/ring_theory/witt_vector/teichmuller.lean index edfd71b3d840c..e6a7e126407f8 100644 --- a/src/ring_theory/witt_vector/teichmuller.lean +++ b/src/ring_theory/witt_vector/teichmuller.lean @@ -9,6 +9,9 @@ import ring_theory.witt_vector.basic /-! # Teichmüller lifts +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines `witt_vector.teichmuller`, a monoid hom `R →* 𝕎 R`, which embeds `r : R` as the `0`-th component of a Witt vector whose other coefficients are `0`. diff --git a/src/topology/algebra/continuous_affine_map.lean b/src/topology/algebra/continuous_affine_map.lean index df8ea46d78521..37500de93756f 100644 --- a/src/topology/algebra/continuous_affine_map.lean +++ b/src/topology/algebra/continuous_affine_map.lean @@ -10,6 +10,9 @@ import topology.algebra.module.basic /-! # Continuous affine maps. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines a type of bundled continuous affine maps. Note that the definition and basic properties established here require minimal assumptions, and do diff --git a/src/topology/algebra/module/character_space.lean b/src/topology/algebra/module/character_space.lean index 1257a49d62540..8de7b152729af 100644 --- a/src/topology/algebra/module/character_space.lean +++ b/src/topology/algebra/module/character_space.lean @@ -11,6 +11,9 @@ import topology.continuous_function.algebra /-! # Character space of a topological algebra +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The character space of a topological algebra is the subset of elements of the weak dual that are also algebra homomorphisms. This space is used in the Gelfand transform, which gives an isomorphism between a commutative C⋆-algebra and continuous functions on the character space diff --git a/src/topology/category/Compactum.lean b/src/topology/category/Compactum.lean index 6bc1fc6992906..b0a26f2fed396 100644 --- a/src/topology/category/Compactum.lean +++ b/src/topology/category/Compactum.lean @@ -15,6 +15,9 @@ import data.set.constructions # Compacta and Compact Hausdorff Spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Recall that, given a monad `M` on `Type*`, an *algebra* for `M` consists of the following data: - A type `X : Type*` - A "structure" map `M X → X`. diff --git a/src/topology/continuous_function/units.lean b/src/topology/continuous_function/units.lean index c88dc320c5ce3..bfa4927cff628 100644 --- a/src/topology/continuous_function/units.lean +++ b/src/topology/continuous_function/units.lean @@ -10,6 +10,9 @@ import topology.continuous_function.algebra /-! # Units of continuous functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file concerns itself with `C(X, M)ˣ` and `C(X, Mˣ)` when `X` is a topological space and `M` has some monoid structure compatible with its topology. -/ From 59150e4a8ab64a88e35d5cfa0a17b762a68acfe9 Mon Sep 17 00:00:00 2001 From: Eric Rodriguez <37984851+ericrbg@users.noreply.github.com> Date: Fri, 9 Jun 2023 07:26:36 +0000 Subject: [PATCH 1161/1291] chore(modular_forms/basic): fix typo in docs (#19166) --- src/number_theory/modular_forms/basic.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/number_theory/modular_forms/basic.lean b/src/number_theory/modular_forms/basic.lean index 08a2663bcdfff..693f1d5696ef8 100644 --- a/src/number_theory/modular_forms/basic.lean +++ b/src/number_theory/modular_forms/basic.lean @@ -56,14 +56,14 @@ structure cusp_form extends slash_invariant_form Γ k := add_decl_doc cusp_form.to_slash_invariant_form /--`modular_form_class F Γ k` says that `F` is a type of bundled functions that extend -`slash_invariant_forms_class` by requiring that the functions be holomorphic and bounded +`slash_invariant_form_class` by requiring that the functions be holomorphic and bounded at infinity. -/ class modular_form_class extends slash_invariant_form_class F Γ k := (holo: ∀ f : F, mdifferentiable 𝓘(ℂ) 𝓘(ℂ) (f : ℍ → ℂ)) (bdd_at_infty : ∀ (f : F) (A : SL(2, ℤ)), is_bounded_at_im_infty (f ∣[k] A)) /--`cusp_form_class F Γ k` says that `F` is a type of bundled functions that extend -`slash_invariant_forms_class` by requiring that the functions be holomorphic and zero +`slash_invariant_form_class` by requiring that the functions be holomorphic and zero at infinity. -/ class cusp_form_class extends slash_invariant_form_class F Γ k := (holo: ∀ f : F, mdifferentiable 𝓘(ℂ) 𝓘(ℂ) (f : ℍ → ℂ)) From cc67cd75b4e54191e13c2e8d722289a89e67e4fa Mon Sep 17 00:00:00 2001 From: Eric Rodriguez Date: Fri, 9 Jun 2023 14:50:20 +0000 Subject: [PATCH 1162/1291] chore(*/centralizer): add forgotten `to_additive`s (#19168) I forgot these in #18861. These are already in the forward-port PR, leanprover-community/mathlib4#4896. --- src/group_theory/subgroup/basic.lean | 5 +++-- src/group_theory/submonoid/centralizer.lean | 6 ++++-- src/group_theory/subsemigroup/centralizer.lean | 8 ++++++-- 3 files changed, 13 insertions(+), 6 deletions(-) diff --git a/src/group_theory/subgroup/basic.lean b/src/group_theory/subgroup/basic.lean index 706259668852c..42878c5673ac4 100644 --- a/src/group_theory/subgroup/basic.lean +++ b/src/group_theory/subgroup/basic.lean @@ -1604,12 +1604,13 @@ set_like.ext' (set.centralizer_univ G) @[to_additive] lemma le_centralizer_iff : H ≤ K.centralizer ↔ K ≤ H.centralizer := ⟨λ h x hx y hy, (h hy x hx).symm, λ h x hx y hy, (h hy x hx).symm⟩ -lemma center_le_centralizer (s) : center G ≤ centralizer s := set.center_subset_centralizer s +@[to_additive] lemma center_le_centralizer (s) : center G ≤ centralizer s := +set.center_subset_centralizer s @[to_additive] lemma centralizer_le (h : H ≤ K) : centralizer K ≤ centralizer H := submonoid.centralizer_le h -@[simp] lemma centralizer_eq_top_iff_subset {s} : centralizer s = ⊤ ↔ s ≤ center G := +@[simp, to_additive] lemma centralizer_eq_top_iff_subset {s} : centralizer s = ⊤ ↔ s ≤ center G := set_like.ext'_iff.trans set.centralizer_eq_top_iff_subset @[to_additive] instance subgroup.centralizer.characteristic [hH : H.characteristic] : diff --git a/src/group_theory/submonoid/centralizer.lean b/src/group_theory/submonoid/centralizer.lean index 47bc7785fb91c..23c2c3fb1e0e5 100644 --- a/src/group_theory/submonoid/centralizer.lean +++ b/src/group_theory/submonoid/centralizer.lean @@ -50,7 +50,8 @@ variables {S} @[to_additive] lemma mem_centralizer_iff {z : M} : z ∈ centralizer S ↔ ∀ g ∈ S, g * z = z * g := iff.rfl -lemma center_le_centralizer (s) : center M ≤ centralizer s := s.center_subset_centralizer +@[to_additive] lemma center_le_centralizer (s) : center M ≤ centralizer s := +s.center_subset_centralizer @[to_additive] instance decidable_mem_centralizer (a) [decidable $ ∀ b ∈ S, b * a = a * b] : decidable (a ∈ centralizer S) := @@ -60,7 +61,8 @@ decidable_of_iff' _ mem_centralizer_iff lemma centralizer_le (h : S ⊆ T) : centralizer T ≤ centralizer S := set.centralizer_subset h -@[simp] lemma centralizer_eq_top_iff_subset {s : set M} : centralizer s = ⊤ ↔ s ⊆ center M := +@[simp, to_additive] lemma centralizer_eq_top_iff_subset {s : set M} : + centralizer s = ⊤ ↔ s ⊆ center M := set_like.ext'_iff.trans set.centralizer_eq_top_iff_subset variables (M) diff --git a/src/group_theory/subsemigroup/centralizer.lean b/src/group_theory/subsemigroup/centralizer.lean index e7dda9535bea3..2d979f0c20cfe 100644 --- a/src/group_theory/subsemigroup/centralizer.lean +++ b/src/group_theory/subsemigroup/centralizer.lean @@ -100,10 +100,12 @@ end lemma centralizer_subset [has_mul M] (h : S ⊆ T) : centralizer T ⊆ centralizer S := λ t ht s hs, ht s (h hs) +@[to_additive add_center_subset_add_centralizer] lemma center_subset_centralizer [has_mul M] (S : set M) : set.center M ⊆ S.centralizer := λ x hx m _, hx m -@[simp] lemma centralizer_eq_top_iff_subset {s : set M} [has_mul M] : +@[simp, to_additive add_centralizer_eq_top_iff_subset] +lemma centralizer_eq_top_iff_subset {s : set M} [has_mul M] : centralizer s = set.univ ↔ s ⊆ center M := eq_top_iff.trans $ ⟨λ h x hx g, (h trivial _ hx).symm, λ h x _ m hm, (h hm x).symm⟩ @@ -144,13 +146,15 @@ iff.rfl decidable (a ∈ centralizer S) := decidable_of_iff' _ mem_centralizer_iff +@[to_additive] lemma center_le_centralizer (S) : center M ≤ centralizer S := S.center_subset_centralizer @[to_additive] lemma centralizer_le (h : S ⊆ T) : centralizer T ≤ centralizer S := set.centralizer_subset h -@[simp] lemma centralizer_eq_top_iff_subset {s : set M} : centralizer s = ⊤ ↔ s ⊆ center M := +@[simp, to_additive] +lemma centralizer_eq_top_iff_subset {s : set M} : centralizer s = ⊤ ↔ s ⊆ center M := set_like.ext'_iff.trans set.centralizer_eq_top_iff_subset variables (M) From ba074af83b6cf54c3104e59402b39410ddbd6dca Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Fri, 9 Jun 2023 17:26:36 +0000 Subject: [PATCH 1163/1291] feat(probability/martingale): optional sampling theorem (#14065) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We prove the optional sampling theorem: if `τ` is a bounded stopping time and `σ` is another stopping time, then the value of a martingale `f` at the stopping time `min τ σ` is almost everywhere equal to `μ[stopped_value f τ | hσ.measurable_space]`. Co-authored-by: RemyDegenne Co-authored-by: Rémy Degenne --- src/probability/martingale/basic.lean | 2 +- .../martingale/optional_sampling.lean | 236 ++++++++++++++++++ src/probability/process/stopping.lean | 70 ++++++ 3 files changed, 307 insertions(+), 1 deletion(-) create mode 100644 src/probability/martingale/optional_sampling.lean diff --git a/src/probability/martingale/basic.lean b/src/probability/martingale/basic.lean index 222656ee7ace4..9501a4443e29f 100644 --- a/src/probability/martingale/basic.lean +++ b/src/probability/martingale/basic.lean @@ -352,7 +352,7 @@ begin refine ⟨hf.1.smul c, λ i j hij, _, λ i, (hf.2.2 i).smul c⟩, refine (condexp_smul c (f j)).le.trans _, filter_upwards [hf.2.1 i j hij] with _ hle, - simp, + simp_rw [pi.smul_apply], exact smul_le_smul_of_nonneg hle hc, end diff --git a/src/probability/martingale/optional_sampling.lean b/src/probability/martingale/optional_sampling.lean new file mode 100644 index 0000000000000..61092ba02e0a4 --- /dev/null +++ b/src/probability/martingale/optional_sampling.lean @@ -0,0 +1,236 @@ +/- +Copyright (c) 2023 Rémy Degenne. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Rémy Degenne +-/ + +import order.succ_pred.linear_locally_finite +import probability.martingale.basic + +/-! +# Optional sampling theorem + +If `τ` is a bounded stopping time and `σ` is another stopping time, then the value of a martingale +`f` at the stopping time `min τ σ` is almost everywhere equal to +`μ[stopped_value f τ | hσ.measurable_space]`. + +## Main results + +* `stopped_value_ae_eq_condexp_of_le_const`: the value of a martingale `f` at a stopping time `τ` + bounded by `n` is the conditional expectation of `f n` with respect to the σ-algebra generated by + `τ`. +* `stopped_value_ae_eq_condexp_of_le`: if `τ` and `σ` are two stopping times with `σ ≤ τ` and `τ` is + bounded, then the value of a martingale `f` at `σ` is the conditional expectation of its value at + `τ` with respect to the σ-algebra generated by `σ`. +* `stopped_value_min_ae_eq_condexp`: the optional sampling theorem. If `τ` is a bounded stopping + time and `σ` is another stopping time, then the value of a martingale `f` at the stopping time + `min τ σ` is almost everywhere equal to the conditional expectation of `f` stopped at `τ` + with respect to the σ-algebra generated by `σ`. + +-/ + +open_locale measure_theory big_operators ennreal +open topological_space + +-- TODO after the port: move to topology/instances/discrete +@[priority 100] +instance discrete_topology.second_countable_topology_of_countable {α : Type*} [topological_space α] + [discrete_topology α] [countable α] : + second_countable_topology α := +@discrete_topology.second_countable_topology_of_encodable _ _ _ (encodable.of_countable _) + +namespace measure_theory + +namespace martingale + +variables {Ω E : Type*} {m : measurable_space Ω} {μ : measure Ω} + [normed_add_comm_group E] [normed_space ℝ E] [complete_space E] + +section first_countable_topology + +variables {ι : Type*} [linear_order ι] [topological_space ι] [order_topology ι] + [first_countable_topology ι] + {ℱ : filtration ι m} [sigma_finite_filtration μ ℱ] {τ σ : Ω → ι} {f : ι → Ω → E} {i n : ι} + +lemma condexp_stopping_time_ae_eq_restrict_eq_const + [(filter.at_top : filter ι).is_countably_generated] + (h : martingale f ℱ μ) (hτ : is_stopping_time ℱ τ) [sigma_finite (μ.trim hτ.measurable_space_le)] + (hin : i ≤ n) : + μ[f n | hτ.measurable_space] =ᵐ[μ.restrict {x | τ x = i}] f i := +begin + refine filter.eventually_eq.trans _ (ae_restrict_of_ae (h.condexp_ae_eq hin)), + refine condexp_ae_eq_restrict_of_measurable_space_eq_on hτ.measurable_space_le (ℱ.le i) + (hτ.measurable_set_eq' i) (λ t, _), + rw [set.inter_comm _ t, is_stopping_time.measurable_set_inter_eq_iff], +end + +lemma condexp_stopping_time_ae_eq_restrict_eq_const_of_le_const + (h : martingale f ℱ μ) (hτ : is_stopping_time ℱ τ) (hτ_le : ∀ x, τ x ≤ n) + [sigma_finite (μ.trim (hτ.measurable_space_le_of_le hτ_le))] (i : ι) : + μ[f n | hτ.measurable_space] =ᵐ[μ.restrict {x | τ x = i}] f i := +begin + by_cases hin : i ≤ n, + { refine filter.eventually_eq.trans _ (ae_restrict_of_ae (h.condexp_ae_eq hin)), + refine condexp_ae_eq_restrict_of_measurable_space_eq_on (hτ.measurable_space_le_of_le hτ_le) + (ℱ.le i) (hτ.measurable_set_eq' i) (λ t, _), + rw [set.inter_comm _ t, is_stopping_time.measurable_set_inter_eq_iff], }, + { suffices : {x : Ω | τ x = i} = ∅, by simp [this], + ext1 x, + simp only [set.mem_set_of_eq, set.mem_empty_iff_false, iff_false], + rintro rfl, + exact hin (hτ_le x), }, +end + +lemma stopped_value_ae_eq_restrict_eq + (h : martingale f ℱ μ) (hτ : is_stopping_time ℱ τ) (hτ_le : ∀ x, τ x ≤ n) + [sigma_finite (μ.trim ((hτ.measurable_space_le_of_le hτ_le)))] (i : ι) : + stopped_value f τ =ᵐ[μ.restrict {x | τ x = i}] μ[f n | hτ.measurable_space] := +begin + refine filter.eventually_eq.trans _ + (condexp_stopping_time_ae_eq_restrict_eq_const_of_le_const h hτ hτ_le i).symm, + rw [filter.eventually_eq, ae_restrict_iff' (ℱ.le _ _ (hτ.measurable_set_eq i))], + refine filter.eventually_of_forall (λ x hx, _), + rw set.mem_set_of_eq at hx, + simp_rw [stopped_value, hx], +end + +/-- The value of a martingale `f` at a stopping time `τ` bounded by `n` is the conditional +expectation of `f n` with respect to the σ-algebra generated by `τ`. -/ +lemma stopped_value_ae_eq_condexp_of_le_const_of_countable_range + (h : martingale f ℱ μ) (hτ : is_stopping_time ℱ τ) + (hτ_le : ∀ x, τ x ≤ n) (h_countable_range : (set.range τ).countable) + [sigma_finite (μ.trim (hτ.measurable_space_le_of_le hτ_le))] : + stopped_value f τ =ᵐ[μ] μ[f n | hτ.measurable_space] := +begin + have : set.univ = ⋃ i ∈ (set.range τ), {x | τ x = i}, + { ext1 x, + simp only [set.mem_univ, set.mem_range, true_and, set.Union_exists, set.Union_Union_eq', + set.mem_Union, set.mem_set_of_eq, exists_apply_eq_apply'], }, + nth_rewrite 0 ← @measure.restrict_univ Ω _ μ, + rw [this, ae_eq_restrict_bUnion_iff _ h_countable_range], + exact λ i hi, stopped_value_ae_eq_restrict_eq h _ hτ_le i, +end + +/-- The value of a martingale `f` at a stopping time `τ` bounded by `n` is the conditional +expectation of `f n` with respect to the σ-algebra generated by `τ`. -/ +lemma stopped_value_ae_eq_condexp_of_le_const [countable ι] + (h : martingale f ℱ μ) (hτ : is_stopping_time ℱ τ) (hτ_le : ∀ x, τ x ≤ n) + [sigma_finite (μ.trim (hτ.measurable_space_le_of_le hτ_le))] : + stopped_value f τ =ᵐ[μ] μ[f n | hτ.measurable_space] := +h.stopped_value_ae_eq_condexp_of_le_const_of_countable_range hτ hτ_le (set.to_countable _) + +/-- If `τ` and `σ` are two stopping times with `σ ≤ τ` and `τ` is bounded, then the value of a +martingale `f` at `σ` is the conditional expectation of its value at `τ` with respect to the +σ-algebra generated by `σ`. -/ +lemma stopped_value_ae_eq_condexp_of_le_of_countable_range + (h : martingale f ℱ μ) (hτ : is_stopping_time ℱ τ) (hσ : is_stopping_time ℱ σ) + (hσ_le_τ : σ ≤ τ) (hτ_le : ∀ x, τ x ≤ n) + (hτ_countable_range : (set.range τ).countable) (hσ_countable_range : (set.range σ).countable) + [sigma_finite (μ.trim (hσ.measurable_space_le_of_le (λ x, (hσ_le_τ x).trans (hτ_le x))))] : + stopped_value f σ =ᵐ[μ] μ[stopped_value f τ | hσ.measurable_space] := +begin + haveI : sigma_finite (μ.trim (hτ.measurable_space_le_of_le hτ_le)), + { exact sigma_finite_trim_mono _ (is_stopping_time.measurable_space_mono hσ hτ hσ_le_τ), }, + have : μ[stopped_value f τ|hσ.measurable_space] + =ᵐ[μ] μ[μ[f n|hτ.measurable_space] | hσ.measurable_space], + from condexp_congr_ae (h.stopped_value_ae_eq_condexp_of_le_const_of_countable_range hτ hτ_le + hτ_countable_range), + refine (filter.eventually_eq.trans _ + (condexp_condexp_of_le _ (hτ.measurable_space_le_of_le hτ_le)).symm).trans this.symm, + { exact h.stopped_value_ae_eq_condexp_of_le_const_of_countable_range hσ + (λ x, (hσ_le_τ x).trans (hτ_le x)) hσ_countable_range, }, + { exact hσ.measurable_space_mono hτ hσ_le_τ, }, +end + +/-- If `τ` and `σ` are two stopping times with `σ ≤ τ` and `τ` is bounded, then the value of a +martingale `f` at `σ` is the conditional expectation of its value at `τ` with respect to the +σ-algebra generated by `σ`. -/ +lemma stopped_value_ae_eq_condexp_of_le [countable ι] + (h : martingale f ℱ μ) (hτ : is_stopping_time ℱ τ) (hσ : is_stopping_time ℱ σ) + (hσ_le_τ : σ ≤ τ) (hτ_le : ∀ x, τ x ≤ n) [sigma_finite (μ.trim hσ.measurable_space_le)] : + stopped_value f σ =ᵐ[μ] μ[stopped_value f τ | hσ.measurable_space] := +h.stopped_value_ae_eq_condexp_of_le_of_countable_range hτ hσ hσ_le_τ hτ_le + (set.to_countable _) (set.to_countable _) + +end first_countable_topology + +section subset_of_nat + +/-! In the following results the index set verifies `[linear_order ι] [locally_finite_order ι]` and +`[order_bot ι]`, which means that it is order-isomorphic to a subset of `ℕ`. `ι` is equipped with +the discrete topology, which is also the order topology, and is a measurable space with the Borel +σ-algebra. -/ + +variables {ι : Type*} [linear_order ι] [locally_finite_order ι] [order_bot ι] + [topological_space ι] [discrete_topology ι] [measurable_space ι] [borel_space ι] + [measurable_space E] [borel_space E] [second_countable_topology E] + {ℱ : filtration ι m} {τ σ : Ω → ι} {f : ι → Ω → E} {i n : ι} + +lemma condexp_stopped_value_stopping_time_ae_eq_restrict_le + (h : martingale f ℱ μ) (hτ : is_stopping_time ℱ τ) (hσ : is_stopping_time ℱ σ) + [sigma_finite (μ.trim hσ.measurable_space_le)] (hτ_le : ∀ x, τ x ≤ n) : + μ[stopped_value f τ | hσ.measurable_space] =ᵐ[μ.restrict {x : Ω | τ x ≤ σ x}] stopped_value f τ := +begin + rw ae_eq_restrict_iff_indicator_ae_eq + (hτ.measurable_space_le _ (hτ.measurable_set_le_stopping_time hσ)), + swap, apply_instance, + refine (condexp_indicator (integrable_stopped_value ι hτ h.integrable hτ_le) + (hτ.measurable_set_stopping_time_le hσ)).symm.trans _, + have h_int : integrable ({ω : Ω | τ ω ≤ σ ω}.indicator (stopped_value (λ (n : ι), f n) τ)) μ, + { refine (integrable_stopped_value ι hτ h.integrable hτ_le).indicator _, + exact hτ.measurable_space_le _ (hτ.measurable_set_le_stopping_time hσ), }, + have h_meas : ae_strongly_measurable' hσ.measurable_space + ({ω : Ω | τ ω ≤ σ ω}.indicator (stopped_value (λ (n : ι), f n) τ)) μ, + { refine strongly_measurable.ae_strongly_measurable' _, + refine strongly_measurable.strongly_measurable_of_measurable_space_le_on + (hτ.measurable_set_le_stopping_time hσ) _ _ _, + { intros t ht, + rw set.inter_comm _ t at ht ⊢, + rw [hτ.measurable_set_inter_le_iff, is_stopping_time.measurable_set_min_iff hτ hσ] at ht, + exact ht.2, }, + { refine strongly_measurable.indicator _ (hτ.measurable_set_le_stopping_time hσ), + refine measurable.strongly_measurable _, + exact measurable_stopped_value h.adapted.prog_measurable_of_discrete hτ, }, + { intros x hx, + simp only [hx, set.indicator_of_not_mem, not_false_iff], }, }, + exact condexp_of_ae_strongly_measurable' hσ.measurable_space_le h_meas h_int, +end + +/-- **Optional Sampling theorem**. If `τ` is a bounded stopping time and `σ` is another stopping +time, then the value of a martingale `f` at the stopping time `min τ σ` is almost everywhere equal +to the conditional expectation of `f` stopped at `τ` with respect to the σ-algebra generated +by `σ`. -/ +lemma stopped_value_min_ae_eq_condexp [sigma_finite_filtration μ ℱ] + (h : martingale f ℱ μ) (hτ : is_stopping_time ℱ τ) (hσ : is_stopping_time ℱ σ) {n : ι} + (hτ_le : ∀ x, τ x ≤ n) [h_sf_min : sigma_finite (μ.trim (hτ.min hσ).measurable_space_le)] : + stopped_value f (λ x, min (σ x) (τ x)) =ᵐ[μ] μ[stopped_value f τ | hσ.measurable_space] := +begin + refine (h.stopped_value_ae_eq_condexp_of_le hτ (hσ.min hτ) (λ x, min_le_right _ _) hτ_le).trans _, + refine ae_of_ae_restrict_of_ae_restrict_compl {x | σ x ≤ τ x} _ _, + { exact condexp_min_stopping_time_ae_eq_restrict_le hσ hτ, }, + { suffices : μ[stopped_value f τ|(hσ.min hτ).measurable_space] + =ᵐ[μ.restrict {x | τ x ≤ σ x}] μ[stopped_value f τ|hσ.measurable_space], + { rw ae_restrict_iff' (hσ.measurable_space_le _ (hσ.measurable_set_le_stopping_time hτ).compl), + rw [filter.eventually_eq, ae_restrict_iff'] at this, + swap, { exact hτ.measurable_space_le _ (hτ.measurable_set_le_stopping_time hσ), }, + filter_upwards [this] with x hx hx_mem, + simp only [set.mem_compl_iff, set.mem_set_of_eq, not_le] at hx_mem, + exact hx hx_mem.le, }, + refine filter.eventually_eq.trans _ + ((condexp_min_stopping_time_ae_eq_restrict_le hτ hσ).trans _), + { exact stopped_value f τ, }, + { rw [is_stopping_time.measurable_space_min, is_stopping_time.measurable_space_min, inf_comm] }, + { have h1 : μ[stopped_value f τ|hτ.measurable_space] = stopped_value f τ, + { refine condexp_of_strongly_measurable hτ.measurable_space_le _ _, + { refine measurable.strongly_measurable _, + exact measurable_stopped_value h.adapted.prog_measurable_of_discrete hτ, }, + { exact integrable_stopped_value ι hτ h.integrable hτ_le, }, }, + rw h1, + exact (condexp_stopped_value_stopping_time_ae_eq_restrict_le h hτ hσ hτ_le).symm, }, }, +end + +end subset_of_nat + +end martingale + +end measure_theory diff --git a/src/probability/process/stopping.lean b/src/probability/process/stopping.lean index 29832258abf99..be11247dd49c5 100644 --- a/src/probability/process/stopping.lean +++ b/src/probability/process/stopping.lean @@ -1216,4 +1216,74 @@ by { ext ω, rw stopped_value, by_cases hx : ω ∈ s; simp [hx] } end piecewise_const +section condexp +/-! ### Conditional expectation with respect to the σ-algebra generated by a stopping time -/ + +variables [linear_order ι] {μ : measure Ω} {ℱ : filtration ι m} {τ σ : Ω → ι} + {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] [complete_space E] {f : Ω → E} + +lemma condexp_stopping_time_ae_eq_restrict_eq_of_countable_range [sigma_finite_filtration μ ℱ] + (hτ : is_stopping_time ℱ τ) (h_countable : (set.range τ).countable) + [sigma_finite (μ.trim (hτ.measurable_space_le_of_countable_range h_countable))] (i : ι) : + μ[f | hτ.measurable_space] =ᵐ[μ.restrict {x | τ x = i}] μ[f | ℱ i] := +begin + refine condexp_ae_eq_restrict_of_measurable_space_eq_on + (hτ.measurable_space_le_of_countable_range h_countable) (ℱ.le i) + (hτ.measurable_set_eq_of_countable_range' h_countable i) (λ t, _), + rw [set.inter_comm _ t, is_stopping_time.measurable_set_inter_eq_iff], +end + +lemma condexp_stopping_time_ae_eq_restrict_eq_of_countable [countable ι] + [sigma_finite_filtration μ ℱ] + (hτ : is_stopping_time ℱ τ) [sigma_finite (μ.trim hτ.measurable_space_le_of_countable)] (i : ι) : + μ[f | hτ.measurable_space] =ᵐ[μ.restrict {x | τ x = i}] μ[f | ℱ i] := +condexp_stopping_time_ae_eq_restrict_eq_of_countable_range hτ (set.to_countable _) i + +variables [(filter.at_top : filter ι).is_countably_generated] + +lemma condexp_min_stopping_time_ae_eq_restrict_le_const (hτ : is_stopping_time ℱ τ) + (i : ι) [sigma_finite (μ.trim (hτ.min_const i).measurable_space_le)] : + μ[f | (hτ.min_const i).measurable_space] + =ᵐ[μ.restrict {x | τ x ≤ i}] μ[f | hτ.measurable_space] := +begin + haveI : sigma_finite (μ.trim hτ.measurable_space_le), + { have h_le : (hτ.min_const i).measurable_space ≤ hτ.measurable_space, + { rw is_stopping_time.measurable_space_min_const, + exact inf_le_left, }, + exact sigma_finite_trim_mono _ h_le, }, + refine (condexp_ae_eq_restrict_of_measurable_space_eq_on hτ.measurable_space_le + (hτ.min_const i).measurable_space_le (hτ.measurable_set_le' i) (λ t, _)).symm, + rw [set.inter_comm _ t, hτ.measurable_set_inter_le_const_iff], +end + +variables [topological_space ι] [order_topology ι] + +lemma condexp_stopping_time_ae_eq_restrict_eq + [first_countable_topology ι] [sigma_finite_filtration μ ℱ] + (hτ : is_stopping_time ℱ τ) [sigma_finite (μ.trim hτ.measurable_space_le)] (i : ι) : + μ[f | hτ.measurable_space] =ᵐ[μ.restrict {x | τ x = i}] μ[f | ℱ i] := +begin + refine condexp_ae_eq_restrict_of_measurable_space_eq_on + hτ.measurable_space_le (ℱ.le i) (hτ.measurable_set_eq' i) (λ t, _), + rw [set.inter_comm _ t, is_stopping_time.measurable_set_inter_eq_iff], +end + +lemma condexp_min_stopping_time_ae_eq_restrict_le [measurable_space ι] + [second_countable_topology ι] [borel_space ι] + (hτ : is_stopping_time ℱ τ) (hσ : is_stopping_time ℱ σ) + [sigma_finite (μ.trim (hτ.min hσ).measurable_space_le)] : + μ[f | (hτ.min hσ).measurable_space] =ᵐ[μ.restrict {x | τ x ≤ σ x}] μ[f | hτ.measurable_space] := +begin + haveI : sigma_finite (μ.trim hτ.measurable_space_le), + { have h_le : (hτ.min hσ).measurable_space ≤ hτ.measurable_space, + { rw is_stopping_time.measurable_space_min, + exact inf_le_left, }, + exact sigma_finite_trim_mono _ h_le, }, + refine (condexp_ae_eq_restrict_of_measurable_space_eq_on hτ.measurable_space_le + (hτ.min hσ).measurable_space_le (hτ.measurable_set_le_stopping_time hσ) (λ t, _)).symm, + rw [set.inter_comm _ t, is_stopping_time.measurable_set_inter_le_iff], +end + +end condexp + end measure_theory From 96aa788f3e443efb3dace8a634258a9259364f43 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 9 Jun 2023 17:26:37 +0000 Subject: [PATCH 1164/1291] =?UTF-8?q?feat(linear=5Falgebra/matrix/nonsingu?= =?UTF-8?q?lar=5Finverse):=202=C3=972=20block=20triangular=20matrices=20ar?= =?UTF-8?q?e=20invertible=20iff=20their=20diagonal=20is=20(#18849)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- .../matrix/schur_complement.lean | 161 ++++++++++++++++++ 1 file changed, 161 insertions(+) diff --git a/src/linear_algebra/matrix/schur_complement.lean b/src/linear_algebra/matrix/schur_complement.lean index da1b1d46d3a3b..440f1fbe308be 100644 --- a/src/linear_algebra/matrix/schur_complement.lean +++ b/src/linear_algebra/matrix/schur_complement.lean @@ -15,6 +15,10 @@ This file proves properties of 2×2 block matrices `[A B; C D]` that relate to t * `matrix.det_from_blocks₁₁`, `matrix.det_from_blocks₂₂`: determinant of a block matrix in terms of the Schur complement. + * `matrix.inv_of_from_blocks_zero₂₁_eq`, `matrix.inv_of_from_blocks_zero₁₂_eq`: the inverse of a + block triangular matrix. + * `matrix.is_unit_from_blocks_zero₂₁`, `matrix.is_unit_from_blocks_zero₁₂`: invertibility of a + block triangular matrix. * `matrix.det_one_add_mul_comm`: the **Weinstein–Aronszajn identity**. * `matrix.schur_complement_pos_semidef_iff` : If a matrix `A` is positive definite, then `[A B; Bᴴ D]` is postive semidefinite if and only if `D - Bᴴ A⁻¹ B` is postive semidefinite. @@ -54,6 +58,163 @@ lemma from_blocks_eq_of_invertible₂₂ equiv.sum_comm_apply, from_blocks_submatrix_sum_swap_sum_swap] using from_blocks_eq_of_invertible₁₁ D C B A +section triangular + +/-! #### Block triangular matrices -/ + +/-- An upper-block-triangular matrix is invertible if its diagonal is. -/ +def from_blocks_zero₂₁_invertible (A : matrix m m α) (B : matrix m n α) (D : matrix n n α) + [invertible A] [invertible D] : invertible (from_blocks A B 0 D) := +invertible_of_left_inverse _ (from_blocks (⅟A) (-(⅟A⬝B⬝⅟D)) 0 (⅟D)) $ + by simp_rw [from_blocks_multiply, matrix.mul_zero, matrix.zero_mul, zero_add, add_zero, + matrix.neg_mul, matrix.inv_of_mul_self, matrix.mul_inv_of_mul_self_cancel, add_right_neg, + from_blocks_one] + +/-- A lower-block-triangular matrix is invertible if its diagonal is. -/ +def from_blocks_zero₁₂_invertible (A : matrix m m α) (C : matrix n m α) (D : matrix n n α) + [invertible A] [invertible D] : invertible (from_blocks A 0 C D) := +invertible_of_left_inverse _ (from_blocks (⅟A) 0 (-(⅟D⬝C⬝⅟A)) (⅟D)) $ + -- a symmetry argument is more work than just copying the proof + by simp_rw [from_blocks_multiply, matrix.mul_zero, matrix.zero_mul, zero_add, add_zero, + matrix.neg_mul, matrix.inv_of_mul_self, matrix.mul_inv_of_mul_self_cancel, add_left_neg, + from_blocks_one] + +lemma inv_of_from_blocks_zero₂₁_eq + (A : matrix m m α) (B : matrix m n α) (D : matrix n n α) + [invertible A] [invertible D] [invertible (from_blocks A B 0 D)] : + ⅟(from_blocks A B 0 D) = from_blocks (⅟A) (-(⅟A⬝B⬝⅟D)) 0 (⅟D) := +begin + letI := from_blocks_zero₂₁_invertible A B D, + convert (rfl : ⅟(from_blocks A B 0 D) = _), +end + +lemma inv_of_from_blocks_zero₁₂_eq + (A : matrix m m α) (C : matrix n m α) (D : matrix n n α) + [invertible A] [invertible D] [invertible (from_blocks A 0 C D)] : + ⅟(from_blocks A 0 C D) = from_blocks (⅟A) 0 (-(⅟D⬝C⬝⅟A)) (⅟D) := +begin + letI := from_blocks_zero₁₂_invertible A C D, + convert (rfl : ⅟(from_blocks A 0 C D) = _), +end + +/-- Both diagonal entries of an invertible upper-block-triangular matrix are invertible (by reading +off the diagonal entries of the inverse). -/ +def invertible_of_from_blocks_zero₂₁_invertible + (A : matrix m m α) (B : matrix m n α) (D : matrix n n α) + [invertible (from_blocks A B 0 D)] : invertible A × invertible D := +{ fst := invertible_of_left_inverse _ (⅟(from_blocks A B 0 D)).to_blocks₁₁ $ begin + have := matrix.inv_of_mul_self (from_blocks A B 0 D), + rw [←from_blocks_to_blocks (⅟(from_blocks A B 0 D)), from_blocks_multiply] at this, + replace := congr_arg matrix.to_blocks₁₁ this, + simpa only [matrix.to_blocks_from_blocks₁₁, matrix.mul_zero, add_zero, ←from_blocks_one] + using this, + end, + snd := invertible_of_right_inverse _ (⅟(from_blocks A B 0 D)).to_blocks₂₂ $ begin + have := matrix.mul_inv_of_self (from_blocks A B 0 D), + rw [←from_blocks_to_blocks (⅟(from_blocks A B 0 D)), from_blocks_multiply] at this, + replace := congr_arg matrix.to_blocks₂₂ this, + simpa only [matrix.to_blocks_from_blocks₂₂, matrix.zero_mul, zero_add, ←from_blocks_one] + using this, + end } + +/-- Both diagonal entries of an invertible lower-block-triangular matrix are invertible (by reading +off the diagonal entries of the inverse). -/ +def invertible_of_from_blocks_zero₁₂_invertible + (A : matrix m m α) (C : matrix n m α) (D : matrix n n α) + [invertible (from_blocks A 0 C D)] : invertible A × invertible D := +{ fst := invertible_of_right_inverse _ (⅟(from_blocks A 0 C D)).to_blocks₁₁ $ begin + have := matrix.mul_inv_of_self (from_blocks A 0 C D), + rw [←from_blocks_to_blocks (⅟(from_blocks A 0 C D)), from_blocks_multiply] at this, + replace := congr_arg matrix.to_blocks₁₁ this, + simpa only [matrix.to_blocks_from_blocks₁₁, matrix.zero_mul, add_zero, ←from_blocks_one] + using this, + end, + snd := invertible_of_left_inverse _ (⅟(from_blocks A 0 C D)).to_blocks₂₂ $ begin + have := matrix.inv_of_mul_self (from_blocks A 0 C D), + rw [←from_blocks_to_blocks (⅟(from_blocks A 0 C D)), from_blocks_multiply] at this, + replace := congr_arg matrix.to_blocks₂₂ this, + simpa only [matrix.to_blocks_from_blocks₂₂, matrix.mul_zero, zero_add, ←from_blocks_one] + using this, + end } + +/-- `invertible_of_from_blocks_zero₂₁_invertible` and `from_blocks_zero₂₁_invertible` form +an equivalence. -/ +def from_blocks_zero₂₁_invertible_equiv (A : matrix m m α) (B : matrix m n α) (D : matrix n n α) : + invertible (from_blocks A B 0 D) ≃ invertible A × invertible D := +{ to_fun := λ _, by exactI invertible_of_from_blocks_zero₂₁_invertible A B D, + inv_fun := λ i, by letI := i.1; letI := i.2; exact from_blocks_zero₂₁_invertible A B D, + left_inv := λ _, subsingleton.elim _ _, + right_inv := λ _, subsingleton.elim _ _ } + +/-- `invertible_of_from_blocks_zero₁₂_invertible` and `from_blocks_zero₁₂_invertible` form +an equivalence. -/ +def from_blocks_zero₁₂_invertible_equiv (A : matrix m m α) (C : matrix n m α) (D : matrix n n α) : + invertible (from_blocks A 0 C D) ≃ invertible A × invertible D := +{ to_fun := λ _, by exactI invertible_of_from_blocks_zero₁₂_invertible A C D, + inv_fun := λ i, by letI := i.1; letI := i.2; exact from_blocks_zero₁₂_invertible A C D, + left_inv := λ _, subsingleton.elim _ _, + right_inv := λ _, subsingleton.elim _ _ } + +/-- An upper block-triangular matrix is invertible iff both elements of its diagonal are. + +This is a propositional form of `matrix.from_blocks_zero₂₁_invertible_equiv`. -/ +@[simp] lemma is_unit_from_blocks_zero₂₁ {A : matrix m m α} {B : matrix m n α} {D : matrix n n α} : + is_unit (from_blocks A B 0 D) ↔ is_unit A ∧ is_unit D := +by simp only [← nonempty_invertible_iff_is_unit, ←nonempty_prod, + (from_blocks_zero₂₁_invertible_equiv _ _ _).nonempty_congr] + +/-- A lower block-triangular matrix is invertible iff both elements of its diagonal are. + +This is a propositional form of `matrix.from_blocks_zero₁₂_invertible_equiv` forms an `iff`. -/ +@[simp] lemma is_unit_from_blocks_zero₁₂ {A : matrix m m α} {C : matrix n m α} {D : matrix n n α} : + is_unit (from_blocks A 0 C D) ↔ is_unit A ∧ is_unit D := +by simp only [← nonempty_invertible_iff_is_unit, ←nonempty_prod, + (from_blocks_zero₁₂_invertible_equiv _ _ _).nonempty_congr] + +/-- An expression for the inverse of an upper block-triangular matrix, when either both elements of +diagonal are invertible, or both are not. -/ +lemma inv_from_blocks_zero₂₁_of_is_unit_iff + (A : matrix m m α) (B : matrix m n α) (D : matrix n n α) + (hAD : is_unit A ↔ is_unit D) : + (from_blocks A B 0 D)⁻¹ = from_blocks A⁻¹ (-(A⁻¹⬝B⬝D⁻¹)) 0 D⁻¹ := +begin + by_cases hA : is_unit A, + { have hD := hAD.mp hA, + casesI hA.nonempty_invertible, + casesI hD.nonempty_invertible, + letI := from_blocks_zero₂₁_invertible A B D, + simp_rw [←inv_of_eq_nonsing_inv, inv_of_from_blocks_zero₂₁_eq] }, + { have hD := hAD.not.mp hA, + have : ¬is_unit (from_blocks A B 0 D) := + is_unit_from_blocks_zero₂₁.not.mpr (not_and'.mpr $ λ _, hA), + simp_rw [nonsing_inv_eq_ring_inverse, + ring.inverse_non_unit _ hA, ring.inverse_non_unit _ hD, ring.inverse_non_unit _ this, + matrix.zero_mul, neg_zero, from_blocks_zero] } +end + +/-- An expression for the inverse of a lower block-triangular matrix, when either both elements of +diagonal are invertible, or both are not. -/ +lemma inv_from_blocks_zero₁₂_of_is_unit_iff + (A : matrix m m α) (C : matrix n m α) (D : matrix n n α) + (hAD : is_unit A ↔ is_unit D) : + (from_blocks A 0 C D)⁻¹ = from_blocks A⁻¹ 0 (-(D⁻¹⬝C⬝A⁻¹)) D⁻¹ := +begin + by_cases hA : is_unit A, + { have hD := hAD.mp hA, + casesI hA.nonempty_invertible, + casesI hD.nonempty_invertible, + letI := from_blocks_zero₁₂_invertible A C D, + simp_rw [←inv_of_eq_nonsing_inv, inv_of_from_blocks_zero₁₂_eq] }, + { have hD := hAD.not.mp hA, + have : ¬is_unit (from_blocks A 0 C D) := + is_unit_from_blocks_zero₁₂.not.mpr (not_and'.mpr $ λ _, hA), + simp_rw [nonsing_inv_eq_ring_inverse, + ring.inverse_non_unit _ hA, ring.inverse_non_unit _ hD, ring.inverse_non_unit _ this, + matrix.zero_mul, neg_zero, from_blocks_zero] } +end + +end triangular + /-! ### Lemmas about `matrix.det` -/ section det From a3e83f0fa4391c8740f7d773a7a9b74e311ae2a3 Mon Sep 17 00:00:00 2001 From: Xavier Roblot <46200072+xroblot@users.noreply.github.com> Date: Fri, 9 Jun 2023 20:02:24 +0000 Subject: [PATCH 1165/1291] feat(algebra/module/zlattice): prove some results about Z-lattices (#18266) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit For a ℤ-lattice `L` given by `L := submodule.span ℤ (set.range b)` where `b` is a basis of finite dimensional real vector space `E`, this PR defines the fundamental domain of `L` and proves that it is a fundamental domain in the sense of `measure_theory.group.fundamental_domain`. It also introduces most of the tools that will be needed to prove that a discrete subgroup of `E` that spans `E` over `ℝ` is a ℤ-lattice, see #18477 --- src/algebra/module/zlattice.lean | 251 +++++++++++++++++++++++++++++++ 1 file changed, 251 insertions(+) create mode 100644 src/algebra/module/zlattice.lean diff --git a/src/algebra/module/zlattice.lean b/src/algebra/module/zlattice.lean new file mode 100644 index 0000000000000..2d33d30ca3ea2 --- /dev/null +++ b/src/algebra/module/zlattice.lean @@ -0,0 +1,251 @@ +/- +Copyright (c) 2023 Xavier Roblot. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Xavier Roblot +-/ +import measure_theory.group.fundamental_domain + +/-! +# ℤ-lattices + +Let `E` be a finite dimensional vector space over a `normed_linear_ordered_field` `K` with a solid +norm and that is also a `floor_ring`, e.g. `ℚ` or `ℝ`. A (full) ℤ-lattice `L` of `E` is a discrete +subgroup of `E` such that `L` spans `E` over `K`. + +The ℤ-lattice `L` can be defined in two ways: +* For `b` a basis of `E`, then `submodule.span ℤ (set.range b)` is a ℤ-lattice of `E`. +* As an `add_subgroup E` with the additional properties: + `∀ r : ℝ, (L ∩ metric.closed_ball 0 r).finite`, that is `L` is discrete + `submodule.span ℝ (L : set E) = ⊤`, that is `L` spans `E` over `K`. + +## Main result +* `zspan.is_add_fundamental_domain`: for a ℤ-lattice `submodule.span ℤ (set.range b)`, proves that +the set defined by `zspan.fundamental_domain` is a fundamental domain. + +-/ + +open_locale big_operators + +noncomputable theory + +namespace zspan + +open measure_theory measurable_set submodule + +variables {E ι : Type*} + +section normed_lattice_field + +variables {K : Type*} [normed_linear_ordered_field K] +variables [normed_add_comm_group E] [normed_space K E] +variables (b : basis ι K E) + +/-- The fundamental domain of the ℤ-lattice spanned by `b`. See `zspan.is_add_fundamental_domain` +for the proof that it is the fundamental domain. -/ +def fundamental_domain : set E := { m | ∀ i, b.repr m i ∈ set.Ico (0 : K) 1 } + +@[simp] +lemma mem_fundamental_domain {m : E} : + m ∈ fundamental_domain b ↔ ∀ i, b.repr m i ∈ set.Ico (0 : K) 1 := iff.rfl + +variables [floor_ring K] + +section fintype + +variable [fintype ι] + +/-- The map that sends a vector of `E` to the element of the ℤ-lattice spanned by `b` obtained +by rounding down its coordinates on the basis `b`. -/ +def floor (m : E) : span ℤ (set.range b) := ∑ i, ⌊b.repr m i⌋ • b.restrict_scalars ℤ i + +/-- The map that sends a vector of `E` to the element of the ℤ-lattice spanned by `b` obtained +by rounding up its coordinates on the basis `b`. -/ +def ceil (m : E) : span ℤ (set.range b) := ∑ i, ⌈b.repr m i⌉ • b.restrict_scalars ℤ i + +@[simp] +lemma repr_floor_apply (m : E) (i : ι) : + b.repr (floor b m) i = ⌊b.repr m i⌋ := +by { classical ; simp only [floor, zsmul_eq_smul_cast K, b.repr.map_smul, finsupp.single_apply, + finset.sum_apply', basis.repr_self, finsupp.smul_single', mul_one, finset.sum_ite_eq', coe_sum, + finset.mem_univ, if_true, coe_smul_of_tower, basis.restrict_scalars_apply, linear_equiv.map_sum] } + +@[simp] +lemma repr_ceil_apply (m : E) (i : ι) : + b.repr (ceil b m) i = ⌈b.repr m i⌉ := +by { classical ; simp only [ceil, zsmul_eq_smul_cast K, b.repr.map_smul, finsupp.single_apply, + finset.sum_apply', basis.repr_self, finsupp.smul_single', mul_one, finset.sum_ite_eq', coe_sum, + finset.mem_univ, if_true, coe_smul_of_tower, basis.restrict_scalars_apply, linear_equiv.map_sum] } + +@[simp] +lemma floor_eq_self_of_mem (m : E) (h : m ∈ span ℤ (set.range b)) : (floor b m : E) = m := +begin + apply b.ext_elem, + simp_rw [repr_floor_apply b], + intro i, + obtain ⟨z, hz⟩ := (b.mem_span_iff_repr_mem ℤ _).mp h i, + rw [← hz], + exact congr_arg (coe : ℤ → K) (int.floor_int_cast z), +end + +@[simp] +lemma ceil_eq_self_of_mem (m : E) (h : m ∈ span ℤ (set.range b)) : (ceil b m : E) = m := +begin + apply b.ext_elem, + simp_rw [repr_ceil_apply b], + intro i, + obtain ⟨z, hz⟩ := (b.mem_span_iff_repr_mem ℤ _).mp h i, + rw [← hz], + exact congr_arg (coe : ℤ → K) (int.ceil_int_cast z), +end + +/-- The map that sends a vector `E` to the fundamental domain of the lattice, +see `zspan.fract_mem_fundamental_domain`. -/ +def fract (m : E) : E := m - floor b m + +lemma fract_apply (m : E) : fract b m = m - floor b m := rfl + +@[simp] +lemma repr_fract_apply (m : E) (i : ι): + b.repr (fract b m) i = int.fract (b.repr m i) := +by rw [fract, map_sub, finsupp.coe_sub, pi.sub_apply, repr_floor_apply, int.fract] + +@[simp] +lemma fract_fract (m : E) : fract b (fract b m) = fract b m := +basis.ext_elem b (λ _, by { classical ; simp only [repr_fract_apply, int.fract_fract] }) + +@[simp] +lemma fract_zspan_add (m : E) {v : E} (h : v ∈ span ℤ (set.range b)) : + fract b (v + m) = fract b m := +begin + classical, + refine (basis.ext_elem_iff b).mpr (λ i, _), + simp_rw [repr_fract_apply, int.fract_eq_fract], + use (b.restrict_scalars ℤ).repr ⟨v, h⟩ i, + rw [map_add, finsupp.coe_add, pi.add_apply, add_tsub_cancel_right, + ← (eq_int_cast (algebra_map ℤ K) _), basis.restrict_scalars_repr_apply, coe_mk], +end + +@[simp] +lemma fract_add_zspan (m : E) {v : E} (h : v ∈ span ℤ (set.range b)) : + fract b (m + v) = fract b m := by { rw [add_comm, fract_zspan_add b m h] } + +variable {b} + +lemma fract_eq_self {x : E} : + fract b x = x ↔ x ∈ fundamental_domain b := +by { classical ; simp only [basis.ext_elem_iff b, repr_fract_apply, int.fract_eq_self, + mem_fundamental_domain, set.mem_Ico] } + +variable (b) + +lemma fract_mem_fundamental_domain (x : E) : + fract b x ∈ fundamental_domain b := fract_eq_self.mp (fract_fract b _) + +lemma fract_eq_fract (m n : E) : + fract b m = fract b n ↔ -m + n ∈ span ℤ (set.range b) := +begin + classical, + rw [eq_comm, basis.ext_elem_iff b], + simp_rw [repr_fract_apply, int.fract_eq_fract, eq_comm, basis.mem_span_iff_repr_mem, + sub_eq_neg_add, map_add, linear_equiv.map_neg, finsupp.coe_add, finsupp.coe_neg, pi.add_apply, + pi.neg_apply, ← (eq_int_cast (algebra_map ℤ K) _), set.mem_range], +end + +lemma norm_fract_le [has_solid_norm K] (m : E) : + ‖fract b m‖ ≤ ∑ i, ‖b i‖ := +begin + classical, + calc + ‖fract b m‖ = ‖∑ i, b.repr (fract b m) i • b i‖ : by rw b.sum_repr + ... = ‖∑ i, int.fract (b.repr m i) • b i‖ : by simp_rw repr_fract_apply + ... ≤ ∑ i, ‖int.fract (b.repr m i) • b i‖ : norm_sum_le _ _ + ... ≤ ∑ i, ‖int.fract (b.repr m i)‖ * ‖b i‖ : by simp_rw norm_smul + ... ≤ ∑ i, ‖b i‖ : finset.sum_le_sum (λ i _, _), + suffices : ‖int.fract (((b.repr) m) i)‖ ≤ 1, + { convert mul_le_mul_of_nonneg_right this (norm_nonneg _ : 0 ≤ ‖b i ‖), + exact (one_mul _).symm, }, + rw (norm_one.symm : 1 = ‖(1 : K)‖), + apply norm_le_norm_of_abs_le_abs, + rw [abs_one, int.abs_fract], + exact le_of_lt (int.fract_lt_one _), +end + +section unique + +variable [unique ι] + +@[simp] lemma coe_floor_self (k : K) : (floor (basis.singleton ι K) k : K) = ⌊k⌋ := +basis.ext_elem _ (λ _, by rw [repr_floor_apply, basis.singleton_repr, basis.singleton_repr]) + +@[simp] lemma coe_fract_self (k : K) : (fract (basis.singleton ι K) k : K) = int.fract k := +basis.ext_elem _ (λ _, by rw [repr_fract_apply, basis.singleton_repr, basis.singleton_repr]) + +end unique + +end fintype + +lemma fundamental_domain_bounded [finite ι] [has_solid_norm K] : + metric.bounded (fundamental_domain b) := +begin + casesI nonempty_fintype ι, + use 2 * ∑ j, ‖b j‖, + intros x hx y hy, + refine le_trans (dist_le_norm_add_norm x y) _, + rw [← fract_eq_self.mpr hx, ← fract_eq_self.mpr hy], + refine (add_le_add (norm_fract_le b x) (norm_fract_le b y)).trans _, + rw ← two_mul, +end + +lemma vadd_mem_fundamental_domain [fintype ι] (y : span ℤ (set.range b)) (x : E) : + y +ᵥ x ∈ fundamental_domain b ↔ y = -floor b x := +by rw [subtype.ext_iff, ← add_right_inj x, add_subgroup_class.coe_neg, ← sub_eq_add_neg, + ← fract_apply, ← fract_zspan_add b _ (subtype.mem y), add_comm, ← vadd_eq_add, ← vadd_def, + eq_comm, ← fract_eq_self] + +lemma exist_unique_vadd_mem_fundamental_domain [finite ι] (x : E) : + ∃! v : span ℤ (set.range b), v +ᵥ x ∈ fundamental_domain b := +begin + casesI nonempty_fintype ι, + refine ⟨-floor b x, _, λ y h, _⟩, + { exact (vadd_mem_fundamental_domain b (-floor b x) x).mpr rfl, }, + { exact (vadd_mem_fundamental_domain b y x).mp h, }, +end + +end normed_lattice_field + +section real + +variables [normed_add_comm_group E] [normed_space ℝ E] +variables (b : basis ι ℝ E) + +@[measurability] +lemma fundamental_domain_measurable_set [measurable_space E] [opens_measurable_space E] + [finite ι] : + measurable_set (fundamental_domain b) := +begin + haveI : finite_dimensional ℝ E := finite_dimensional.of_fintype_basis b, + let f := (finsupp.linear_equiv_fun_on_finite ℝ ℝ ι).to_linear_map.comp b.repr.to_linear_map, + let D : set (ι → ℝ) := set.pi set.univ (λ i : ι, (set.Ico (0 : ℝ) 1)), + rw ( _ : fundamental_domain b = f⁻¹' D), + { refine measurable_set_preimage (linear_map.continuous_of_finite_dimensional f).measurable _, + exact pi set.univ.to_countable (λ (i : ι) (H : i ∈ set.univ), measurable_set_Ico), }, + { ext, + simp only [fundamental_domain, set.mem_set_of_eq, linear_map.coe_comp, + linear_equiv.coe_to_linear_map, set.mem_preimage, function.comp_app, set.mem_univ_pi, + finsupp.linear_equiv_fun_on_finite_apply], }, +end + +/-- For a ℤ-lattice `submodule.span ℤ (set.range b)`, proves that the set defined +by `zspan.fundamental_domain` is a fundamental domain. -/ +protected lemma is_add_fundamental_domain [finite ι] [measurable_space E] + [opens_measurable_space E] (μ : measure E) : + is_add_fundamental_domain (span ℤ (set.range b)).to_add_subgroup (fundamental_domain b) μ := +begin + casesI nonempty_fintype ι, + exact is_add_fundamental_domain.mk' (null_measurable_set (fundamental_domain_measurable_set b)) + (λ x, exist_unique_vadd_mem_fundamental_domain b x), +end + +end real + +end zspan From fd4551cfe4b7484b81c2c9ba3405edae27659676 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Sat, 10 Jun 2023 05:31:39 +0000 Subject: [PATCH 1166/1291] chore(*): add mathlib4 synchronization comments (#19171) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.lie.cartan_subalgebra` * `algebra.lie.nilpotent` * `analysis.ODE.picard_lindelof` * `analysis.calculus.affine_map` * `analysis.complex.abs_max` * `analysis.complex.cauchy_integral` * `analysis.complex.liouville` * `analysis.complex.locally_uniform_limit` * `analysis.complex.phragmen_lindelof` * `analysis.complex.removable_singularity` * `analysis.complex.schwarz` * `analysis.convex.measure` * `analysis.fourier.fourier_transform` * `analysis.normed_space.add_torsor_bases` * `analysis.normed_space.continuous_affine_map` * `analysis.special_functions.improper_integrals` * `analysis.special_functions.stirling` * `analysis.special_functions.trigonometric.euler_sine_prod` * `analysis.sum_integral_comparisons` * `category_theory.closed.cartesian` * `category_theory.closed.types` * `category_theory.closed.zero` * `category_theory.monoidal.Mod_` * `data.real.pi.bounds` * `data.vector3` * `field_theory.fixed` * `field_theory.is_alg_closed.basic` * `field_theory.normal` * `measure_theory.group.geometry_of_numbers` * `measure_theory.group.integration` * `measure_theory.integral.circle_integral` * `measure_theory.integral.circle_transform` * `measure_theory.integral.exp_decay` * `measure_theory.integral.integral_eq_improper` * `measure_theory.integral.layercake` * `number_theory.bernoulli_polynomials` * `probability.kernel.basic` * `topology.category.Profinite.cofiltered_limit` --- src/algebra/lie/cartan_subalgebra.lean | 3 +++ src/algebra/lie/nilpotent.lean | 3 +++ src/analysis/ODE/picard_lindelof.lean | 3 +++ src/analysis/calculus/affine_map.lean | 3 +++ src/analysis/complex/abs_max.lean | 3 +++ src/analysis/complex/cauchy_integral.lean | 3 +++ src/analysis/complex/liouville.lean | 3 +++ src/analysis/complex/locally_uniform_limit.lean | 3 +++ src/analysis/complex/phragmen_lindelof.lean | 3 +++ src/analysis/complex/removable_singularity.lean | 3 +++ src/analysis/complex/schwarz.lean | 3 +++ src/analysis/convex/measure.lean | 3 +++ src/analysis/fourier/fourier_transform.lean | 3 +++ src/analysis/normed_space/add_torsor_bases.lean | 3 +++ src/analysis/normed_space/continuous_affine_map.lean | 3 +++ src/analysis/special_functions/improper_integrals.lean | 3 +++ src/analysis/special_functions/stirling.lean | 3 +++ .../special_functions/trigonometric/euler_sine_prod.lean | 3 +++ src/analysis/sum_integral_comparisons.lean | 3 +++ src/category_theory/closed/cartesian.lean | 3 +++ src/category_theory/closed/types.lean | 3 +++ src/category_theory/closed/zero.lean | 3 +++ src/category_theory/monoidal/Mod_.lean | 3 +++ src/data/real/pi/bounds.lean | 3 +++ src/data/vector3.lean | 3 +++ src/field_theory/fixed.lean | 3 +++ src/field_theory/is_alg_closed/basic.lean | 3 +++ src/field_theory/normal.lean | 3 +++ src/measure_theory/group/geometry_of_numbers.lean | 3 +++ src/measure_theory/group/integration.lean | 3 +++ src/measure_theory/integral/circle_integral.lean | 3 +++ src/measure_theory/integral/circle_transform.lean | 3 +++ src/measure_theory/integral/exp_decay.lean | 3 +++ src/measure_theory/integral/integral_eq_improper.lean | 3 +++ src/measure_theory/integral/layercake.lean | 3 +++ src/number_theory/bernoulli_polynomials.lean | 3 +++ src/probability/kernel/basic.lean | 3 +++ src/topology/category/Profinite/cofiltered_limit.lean | 3 +++ 38 files changed, 114 insertions(+) diff --git a/src/algebra/lie/cartan_subalgebra.lean b/src/algebra/lie/cartan_subalgebra.lean index 6ef9ca921b8ff..460f04376cb0b 100644 --- a/src/algebra/lie/cartan_subalgebra.lean +++ b/src/algebra/lie/cartan_subalgebra.lean @@ -9,6 +9,9 @@ import algebra.lie.normalizer /-! # Cartan subalgebras +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Cartan subalgebras are one of the most important concepts in Lie theory. We define them here. The standard example is the set of diagonal matrices in the Lie algebra of matrices. diff --git a/src/algebra/lie/nilpotent.lean b/src/algebra/lie/nilpotent.lean index 1f614879622e9..6d4b290ac1816 100644 --- a/src/algebra/lie/nilpotent.lean +++ b/src/algebra/lie/nilpotent.lean @@ -12,6 +12,9 @@ import ring_theory.nilpotent /-! # Nilpotent Lie algebras +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Like groups, Lie algebras admit a natural concept of nilpotency. More generally, any Lie module carries a natural concept of nilpotency. We define these here via the lower central series. diff --git a/src/analysis/ODE/picard_lindelof.lean b/src/analysis/ODE/picard_lindelof.lean index 9d470e9397d30..2b247eaacb36b 100644 --- a/src/analysis/ODE/picard_lindelof.lean +++ b/src/analysis/ODE/picard_lindelof.lean @@ -9,6 +9,9 @@ import topology.metric_space.contracting /-! # Picard-Lindelöf (Cauchy-Lipschitz) Theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove that an ordinary differential equation $\dot x=v(t, x)$ such that $v$ is Lipschitz continuous in $x$ and continuous in $t$ has a local solution, see `exists_forall_deriv_within_Icc_eq_of_is_picard_lindelof`. diff --git a/src/analysis/calculus/affine_map.lean b/src/analysis/calculus/affine_map.lean index 97f2e2a756d3a..1d7e83abaa361 100644 --- a/src/analysis/calculus/affine_map.lean +++ b/src/analysis/calculus/affine_map.lean @@ -9,6 +9,9 @@ import analysis.calculus.cont_diff /-! # Smooth affine maps +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains results about smoothness of affine maps. ## Main definitions: diff --git a/src/analysis/complex/abs_max.lean b/src/analysis/complex/abs_max.lean index d99498d5098f4..5690a779be860 100644 --- a/src/analysis/complex/abs_max.lean +++ b/src/analysis/complex/abs_max.lean @@ -11,6 +11,9 @@ import topology.algebra.order.extr_closure /-! # Maximum modulus principle +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove several versions of the maximum modulus principle. There are several statements that can be called "the maximum modulus principle" for maps between normed complex spaces. They differ by assumptions on the domain (any space, a nontrivial space, a finite diff --git a/src/analysis/complex/cauchy_integral.lean b/src/analysis/complex/cauchy_integral.lean index 18a80345abc2e..7e49ef907682b 100644 --- a/src/analysis/complex/cauchy_integral.lean +++ b/src/analysis/complex/cauchy_integral.lean @@ -15,6 +15,9 @@ import data.real.cardinality /-! # Cauchy integral formula +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove the Cauchy-Goursat theorem and the Cauchy integral formula for integrals over circles. Most results are formulated for a function `f : ℂ → E` that takes values in a complex Banach space with second countable topology. diff --git a/src/analysis/complex/liouville.lean b/src/analysis/complex/liouville.lean index e297de7e96bd2..aa2b6a8ae40eb 100644 --- a/src/analysis/complex/liouville.lean +++ b/src/analysis/complex/liouville.lean @@ -10,6 +10,9 @@ import analysis.normed_space.completion /-! # Liouville's theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove Liouville's theorem: if `f : E → F` is complex differentiable on the whole space and its range is bounded, then the function is a constant. Various versions of this theorem are formalized in `differentiable.apply_eq_apply_of_bounded`, diff --git a/src/analysis/complex/locally_uniform_limit.lean b/src/analysis/complex/locally_uniform_limit.lean index 26f8e16d780c0..cc91ebd185d1f 100644 --- a/src/analysis/complex/locally_uniform_limit.lean +++ b/src/analysis/complex/locally_uniform_limit.lean @@ -9,6 +9,9 @@ import analysis.calculus.series /-! # Locally uniform limits of holomorphic functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file gathers some results about locally uniform limits of holomorphic functions on an open subset of the complex plane. diff --git a/src/analysis/complex/phragmen_lindelof.lean b/src/analysis/complex/phragmen_lindelof.lean index 71e462a2b384f..770b81587ac39 100644 --- a/src/analysis/complex/phragmen_lindelof.lean +++ b/src/analysis/complex/phragmen_lindelof.lean @@ -9,6 +9,9 @@ import analysis.asymptotics.superpolynomial_decay /-! # Phragmen-Lindelöf principle +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove several versions of the Phragmen-Lindelöf principle, a version of the maximum modulus principle for an unbounded domain. diff --git a/src/analysis/complex/removable_singularity.lean b/src/analysis/complex/removable_singularity.lean index 2e52005a61a26..784b50e998e6d 100644 --- a/src/analysis/complex/removable_singularity.lean +++ b/src/analysis/complex/removable_singularity.lean @@ -10,6 +10,9 @@ import analysis.complex.cauchy_integral /-! # Removable singularity theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove Riemann's removable singularity theorem: if `f : ℂ → E` is complex differentiable in a punctured neighborhood of a point `c` and is bounded in a punctured neighborhood of `c` (or, more generally, $f(z) - f(c)=o((z-c)^{-1})$), then it has a limit at `c` and the diff --git a/src/analysis/complex/schwarz.lean b/src/analysis/complex/schwarz.lean index 4e72bd775073a..3a51789fa3216 100644 --- a/src/analysis/complex/schwarz.lean +++ b/src/analysis/complex/schwarz.lean @@ -9,6 +9,9 @@ import analysis.complex.removable_singularity /-! # Schwarz lemma +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove several versions of the Schwarz lemma. * `complex.norm_deriv_le_div_of_maps_to_ball`, `complex.abs_deriv_le_div_of_maps_to_ball`: if diff --git a/src/analysis/convex/measure.lean b/src/analysis/convex/measure.lean index b4e530389b847..e80f6a9481d9c 100644 --- a/src/analysis/convex/measure.lean +++ b/src/analysis/convex/measure.lean @@ -10,6 +10,9 @@ import measure_theory.measure.lebesgue.eq_haar /-! # Convex sets are null-measurable +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Let `E` be a finite dimensional real vector space, let `μ` be a Haar measure on `E`, let `s` be a convex set in `E`. Then the frontier of `s` has measure zero (see `convex.add_haar_frontier`), hence `s` is a `measure_theory.null_measurable_set` (see `convex.null_measurable_set`). diff --git a/src/analysis/fourier/fourier_transform.lean b/src/analysis/fourier/fourier_transform.lean index c843af425d2a9..99c4df079aeed 100644 --- a/src/analysis/fourier/fourier_transform.lean +++ b/src/analysis/fourier/fourier_transform.lean @@ -11,6 +11,9 @@ import measure_theory.measure.haar.of_basis /-! # The Fourier transform +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We set up the Fourier transform for complex-valued functions on finite-dimensional spaces. ## Design choices diff --git a/src/analysis/normed_space/add_torsor_bases.lean b/src/analysis/normed_space/add_torsor_bases.lean index 6da2223d09839..5e162d8f02c82 100644 --- a/src/analysis/normed_space/add_torsor_bases.lean +++ b/src/analysis/normed_space/add_torsor_bases.lean @@ -11,6 +11,9 @@ import linear_algebra.affine_space.finite_dimensional /-! # Bases in normed affine spaces. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains results about bases in normed affine spaces. ## Main definitions: diff --git a/src/analysis/normed_space/continuous_affine_map.lean b/src/analysis/normed_space/continuous_affine_map.lean index 61513ff3a0f20..ab2ff1080e777 100644 --- a/src/analysis/normed_space/continuous_affine_map.lean +++ b/src/analysis/normed_space/continuous_affine_map.lean @@ -10,6 +10,9 @@ import analysis.normed_space.operator_norm /-! # Continuous affine maps between normed spaces. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file develops the theory of continuous affine maps between affine spaces modelled on normed spaces. diff --git a/src/analysis/special_functions/improper_integrals.lean b/src/analysis/special_functions/improper_integrals.lean index fb45ccfb72a8f..89aa5354e61ab 100644 --- a/src/analysis/special_functions/improper_integrals.lean +++ b/src/analysis/special_functions/improper_integrals.lean @@ -12,6 +12,9 @@ import measure_theory.measure.lebesgue.integral /-! # Evaluation of specific improper integrals +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains some integrability results, and evaluations of integrals, over `ℝ` or over half-infinite intervals in `ℝ`. diff --git a/src/analysis/special_functions/stirling.lean b/src/analysis/special_functions/stirling.lean index 9852179666596..488a468a7c81a 100644 --- a/src/analysis/special_functions/stirling.lean +++ b/src/analysis/special_functions/stirling.lean @@ -9,6 +9,9 @@ import data.real.pi.wallis /-! # Stirling's formula +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves Stirling's formula for the factorial. It states that $n!$ grows asymptotically like $\sqrt{2\pi n}(\frac{n}{e})^n$. diff --git a/src/analysis/special_functions/trigonometric/euler_sine_prod.lean b/src/analysis/special_functions/trigonometric/euler_sine_prod.lean index 7fd83358cee89..955b1d130e7e2 100644 --- a/src/analysis/special_functions/trigonometric/euler_sine_prod.lean +++ b/src/analysis/special_functions/trigonometric/euler_sine_prod.lean @@ -8,6 +8,9 @@ import measure_theory.integral.peak_function /-! # Euler's infinite product for the sine function +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves the infinite product formula $$ \sin \pi z = \pi z \prod_{n = 1}^\infty \left(1 - \frac{z ^ 2}{n ^ 2}\right) $$ diff --git a/src/analysis/sum_integral_comparisons.lean b/src/analysis/sum_integral_comparisons.lean index 611b475c6490d..10380facc6359 100644 --- a/src/analysis/sum_integral_comparisons.lean +++ b/src/analysis/sum_integral_comparisons.lean @@ -10,6 +10,9 @@ import analysis.special_functions.integrals /-! # Comparing sums and integrals +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Summary It is often the case that error terms in analysis can be computed by comparing diff --git a/src/category_theory/closed/cartesian.lean b/src/category_theory/closed/cartesian.lean index 848aeeb9213a2..c561505054da7 100644 --- a/src/category_theory/closed/cartesian.lean +++ b/src/category_theory/closed/cartesian.lean @@ -15,6 +15,9 @@ import category_theory.closed.monoidal /-! # Cartesian closed categories +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given a category with finite products, the cartesian monoidal structure is provided by the local instance `monoidal_of_has_finite_products`. diff --git a/src/category_theory/closed/types.lean b/src/category_theory/closed/types.lean index 30614de6280de..13cb03e263bf3 100644 --- a/src/category_theory/closed/types.lean +++ b/src/category_theory/closed/types.lean @@ -11,6 +11,9 @@ import category_theory.closed.cartesian /-! # Cartesian closure of Type +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Show that `Type u₁` is cartesian closed, and `C ⥤ Type u₁` is cartesian closed for `C` a small category in `Type u₁`. Note this implies that the category of presheaves on a small category `C` is cartesian closed. diff --git a/src/category_theory/closed/zero.lean b/src/category_theory/closed/zero.lean index 720938c6f74b6..73fbe508f9cc5 100644 --- a/src/category_theory/closed/zero.lean +++ b/src/category_theory/closed/zero.lean @@ -12,6 +12,9 @@ import category_theory.limits.shapes.zero_objects /-! # A cartesian closed category with zero object is trivial +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A cartesian closed category with zero object is trivial: it is equivalent to the category with one object and one morphism. diff --git a/src/category_theory/monoidal/Mod_.lean b/src/category_theory/monoidal/Mod_.lean index e88e0ba337171..d3f89f976018c 100644 --- a/src/category_theory/monoidal/Mod_.lean +++ b/src/category_theory/monoidal/Mod_.lean @@ -7,6 +7,9 @@ import category_theory.monoidal.Mon_ /-! # The category of module objects over a monoid object. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ universes v₁ v₂ u₁ u₂ diff --git a/src/data/real/pi/bounds.lean b/src/data/real/pi/bounds.lean index e6e665dd2687d..2a6c704d82b90 100644 --- a/src/data/real/pi/bounds.lean +++ b/src/data/real/pi/bounds.lean @@ -8,6 +8,9 @@ import analysis.special_functions.trigonometric.bounds /-! # Pi +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains lemmas which establish bounds on `real.pi`. Notably, these include `pi_gt_sqrt_two_add_series` and `pi_lt_sqrt_two_add_series`, which bound `π` using series; diff --git a/src/data/vector3.lean b/src/data/vector3.lean index 925a9cb462d77..ac86ca9953c38 100644 --- a/src/data/vector3.lean +++ b/src/data/vector3.lean @@ -9,6 +9,9 @@ import tactic.localized /-! # Alternate definition of `vector` in terms of `fin2` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file provides a locale `vector3` which overrides the `[a, b, c]` notation to create a `vector3` instead of a `list`. diff --git a/src/field_theory/fixed.lean b/src/field_theory/fixed.lean index 115c4b0289428..eb7d527720fe8 100644 --- a/src/field_theory/fixed.lean +++ b/src/field_theory/fixed.lean @@ -13,6 +13,9 @@ import field_theory.tower /-! # Fixed field under a group action. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This is the basis of the Fundamental Theorem of Galois Theory. Given a (finite) group `G` that acts on a field `F`, we define `fixed_points G F`, the subfield consisting of elements of `F` fixed_points by every element of `G`. diff --git a/src/field_theory/is_alg_closed/basic.lean b/src/field_theory/is_alg_closed/basic.lean index 9765b18916db8..6c87fd98682f3 100644 --- a/src/field_theory/is_alg_closed/basic.lean +++ b/src/field_theory/is_alg_closed/basic.lean @@ -11,6 +11,9 @@ import ring_theory.localization.integral /-! # Algebraically Closed Field +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define the typeclass for algebraically closed fields and algebraic closures, and prove some of their properties. diff --git a/src/field_theory/normal.lean b/src/field_theory/normal.lean index c4b959da912e9..6b43eea26977f 100644 --- a/src/field_theory/normal.lean +++ b/src/field_theory/normal.lean @@ -12,6 +12,9 @@ import ring_theory.power_basis /-! # Normal field extensions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define normal field extensions and prove that for a finite extension, being normal is the same as being a splitting field (`normal.of_is_splitting_field` and `normal.exists_is_splitting_field`). diff --git a/src/measure_theory/group/geometry_of_numbers.lean b/src/measure_theory/group/geometry_of_numbers.lean index 57d4ce1720ec1..63dea545779a2 100644 --- a/src/measure_theory/group/geometry_of_numbers.lean +++ b/src/measure_theory/group/geometry_of_numbers.lean @@ -10,6 +10,9 @@ import measure_theory.measure.lebesgue.eq_haar /-! # Geometry of numbers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove some of the fundamental theorems in the geometry of numbers, as studied by Hermann Minkowski. diff --git a/src/measure_theory/group/integration.lean b/src/measure_theory/group/integration.lean index 04d574f004c59..258cf46be347d 100644 --- a/src/measure_theory/group/integration.lean +++ b/src/measure_theory/group/integration.lean @@ -10,6 +10,9 @@ import measure_theory.group.action /-! # Integration on Groups +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We develop properties of integrals with a group as domain. This file contains properties about integrability, Lebesgue integration and Bochner integration. -/ diff --git a/src/measure_theory/integral/circle_integral.lean b/src/measure_theory/integral/circle_integral.lean index 80470acd02a2f..3f553cf5c911d 100644 --- a/src/measure_theory/integral/circle_integral.lean +++ b/src/measure_theory/integral/circle_integral.lean @@ -12,6 +12,9 @@ import analysis.analytic.basic /-! # Integral over a circle in `ℂ` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define `∮ z in C(c, R), f z` to be the integral $\oint_{|z-c|=|R|} f(z)\,dz$ and prove some properties of this integral. We give definition and prove most lemmas for a function `f : ℂ → E`, where `E` is a complex Banach space. For this reason, diff --git a/src/measure_theory/integral/circle_transform.lean b/src/measure_theory/integral/circle_transform.lean index e721eca6870d0..4ba4ce3582103 100644 --- a/src/measure_theory/integral/circle_transform.lean +++ b/src/measure_theory/integral/circle_transform.lean @@ -8,6 +8,9 @@ import measure_theory.integral.circle_integral /-! # Circle integral transform +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define the circle integral transform of a function `f` with complex domain. This is defined as $(2πi)^{-1}\frac{f(x)}{x-w}$ where `x` moves along a circle. We then prove some basic facts about these functions. diff --git a/src/measure_theory/integral/exp_decay.lean b/src/measure_theory/integral/exp_decay.lean index 2090e724c7a2d..aaf1b899783dd 100644 --- a/src/measure_theory/integral/exp_decay.lean +++ b/src/measure_theory/integral/exp_decay.lean @@ -9,6 +9,9 @@ import measure_theory.integral.integral_eq_improper /-! # Integrals with exponential decay at ∞ +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + As easy special cases of general theorems in the library, we prove the following test for integrability: diff --git a/src/measure_theory/integral/integral_eq_improper.lean b/src/measure_theory/integral/integral_eq_improper.lean index 7268b5289486b..1cd7f2b271541 100644 --- a/src/measure_theory/integral/integral_eq_improper.lean +++ b/src/measure_theory/integral/integral_eq_improper.lean @@ -12,6 +12,9 @@ import measure_theory.measure.haar.normed_space /-! # Links between an integral and its "improper" version +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In its current state, mathlib only knows how to talk about definite ("proper") integrals, in the sense that it treats integrals over `[x, +∞)` the same as it treats integrals over `[y, z]`. For example, the integral over `[1, +∞)` is **not** defined to be the limit of diff --git a/src/measure_theory/integral/layercake.lean b/src/measure_theory/integral/layercake.lean index 6cc9602fe672a..058d48d0699f9 100644 --- a/src/measure_theory/integral/layercake.lean +++ b/src/measure_theory/integral/layercake.lean @@ -9,6 +9,9 @@ import analysis.special_functions.integrals /-! # The layer cake formula / Cavalieri's principle / tail probability formula +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove the following layer cake formula. Consider a non-negative measurable function `f` on a sigma-finite measure space. Apply pointwise diff --git a/src/number_theory/bernoulli_polynomials.lean b/src/number_theory/bernoulli_polynomials.lean index 3e7801f402ce0..82df7a3f2e5bd 100644 --- a/src/number_theory/bernoulli_polynomials.lean +++ b/src/number_theory/bernoulli_polynomials.lean @@ -11,6 +11,9 @@ import number_theory.bernoulli /-! # Bernoulli polynomials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The [Bernoulli polynomials](https://en.wikipedia.org/wiki/Bernoulli_polynomials) are an important tool obtained from Bernoulli numbers. diff --git a/src/probability/kernel/basic.lean b/src/probability/kernel/basic.lean index 9bef08cbd895f..26f606148e1c5 100644 --- a/src/probability/kernel/basic.lean +++ b/src/probability/kernel/basic.lean @@ -9,6 +9,9 @@ import measure_theory.constructions.prod.basic /-! # Markov Kernels +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A kernel from a measurable space `α` to another measurable space `β` is a measurable map `α → measure β`, where the measurable space instance on `measure β` is the one defined in `measure_theory.measure.measurable_space`. That is, a kernel `κ` verifies that for all measurable diff --git a/src/topology/category/Profinite/cofiltered_limit.lean b/src/topology/category/Profinite/cofiltered_limit.lean index 1fb982d1ad7d3..0e36712aa5c79 100644 --- a/src/topology/category/Profinite/cofiltered_limit.lean +++ b/src/topology/category/Profinite/cofiltered_limit.lean @@ -12,6 +12,9 @@ import topology.category.Top.limits.konig /-! # Cofiltered limits of profinite sets. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains some theorems about cofiltered limits of profinite sets. ## Main Results From e1e7190efdcefc925cb36f257a8362ef22944204 Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Sat, 10 Jun 2023 17:01:39 +0000 Subject: [PATCH 1167/1291] feat(data/polynomial/div): add X_sub_C_dvd_sub_C_eval (#19120) --- src/data/polynomial/div.lean | 17 ++++++++++++++++- 1 file changed, 16 insertions(+), 1 deletion(-) diff --git a/src/data/polynomial/div.lean b/src/data/polynomial/div.lean index 9abd9e5cf5e15..414337e97976b 100644 --- a/src/data/polynomial/div.lean +++ b/src/data/polynomial/div.lean @@ -395,11 +395,26 @@ lemma mul_div_by_monic_eq_iff_is_root : (X - C a) * (p /ₘ (X - C a)) = p ↔ i by conv {to_rhs, rw ← mod_by_monic_add_div p (monic_X_sub_C a)}; rw [mod_by_monic_X_sub_C_eq_C_eval, h, C_0, zero_add]⟩ -lemma dvd_iff_is_root : (X - C a) ∣ p ↔ is_root p a := +lemma dvd_iff_is_root : X - C a ∣ p ↔ is_root p a := ⟨λ h, by rwa [← dvd_iff_mod_by_monic_eq_zero (monic_X_sub_C _), mod_by_monic_X_sub_C_eq_C_eval, ← C_0, C_inj] at h, λ h, ⟨(p /ₘ (X - C a)), by rw mul_div_by_monic_eq_iff_is_root.2 h⟩⟩ +lemma X_sub_C_dvd_sub_C_eval : X - C a ∣ p - C (p.eval a) := +by rw [dvd_iff_is_root, is_root, eval_sub, eval_C, sub_self] + +lemma mem_span_C_X_sub_C_X_sub_C_iff_eval_eval_eq_zero {b : R[X]} {P : R[X][X]} : + P ∈ (ideal.span {C (X - C a), X - C b} : ideal R[X][X]) ↔ (P.eval b).eval a = 0 := +begin + rw [ideal.mem_span_pair], + split; intro h, + { rcases h with ⟨_, _, rfl⟩, + simp only [eval_C, eval_X, eval_add, eval_sub, eval_mul, add_zero, mul_zero, sub_self] }, + { cases dvd_iff_is_root.mpr h with p hp, + cases @X_sub_C_dvd_sub_C_eval _ b _ P with q hq, + exact ⟨C p, q, by rw [mul_comm, mul_comm q, eq_add_of_sub_eq' hq, hp, C_mul]⟩ } +end + lemma mod_by_monic_X (p : R[X]) : p %ₘ X = C (p.eval 0) := by rw [← mod_by_monic_X_sub_C_eq_C_eval, C_0, sub_zero] From 90b0d53ee6ffa910e5c2a977ce7e2fc704647974 Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Sat, 10 Jun 2023 18:56:02 +0000 Subject: [PATCH 1168/1291] feat(linear_algebra/free_module/ideal_quotient): prove dimension of quotient by ideal equals degree of norm of generator (#19121) Also refactor file into `namespace ideal`. Co-authored-by: Junyan Xu --- .../free_module/ideal_quotient.lean | 44 ++++++----- src/linear_algebra/free_module/norm.lean | 75 +++++++++++++++++++ 2 files changed, 100 insertions(+), 19 deletions(-) create mode 100644 src/linear_algebra/free_module/norm.lean diff --git a/src/linear_algebra/free_module/ideal_quotient.lean b/src/linear_algebra/free_module/ideal_quotient.lean index f0dd7f2c04ceb..e41849fa71da7 100644 --- a/src/linear_algebra/free_module/ideal_quotient.lean +++ b/src/linear_algebra/free_module/ideal_quotient.lean @@ -22,17 +22,20 @@ import linear_algebra.quotient_pi -/ -open_locale big_operators direct_sum +namespace ideal -variables {ι R S : Type*} [comm_ring R] [comm_ring S] [algebra R S] -variables [is_domain R] [is_principal_ideal_ring R] [is_domain S] [fintype ι] +open_locale big_operators direct_sum polynomial + +variables {R S ι : Type*} [comm_ring R] [comm_ring S] [algebra R S] +variables [is_domain R] [is_principal_ideal_ring R] [is_domain S] [finite ι] /-- We can write the quotient of an ideal over a PID as a product of quotients by principal ideals. -/ -noncomputable def ideal.quotient_equiv_pi_span +noncomputable def quotient_equiv_pi_span (I : ideal S) (b : basis ι R S) (hI : I ≠ ⊥) : - (S ⧸ I) ≃ₗ[R] Π i, (R ⧸ ideal.span ({I.smith_coeffs b hI i} : set R)) := + (S ⧸ I) ≃ₗ[R] Π i, (R ⧸ span ({I.smith_coeffs b hI i} : set R)) := begin + haveI := fintype.of_finite ι, -- Choose `e : S ≃ₗ I` and a basis `b'` for `S` that turns the map -- `f := ((submodule.subtype I).restrict_scalars R).comp e` into a diagonal matrix: -- there is an `a : ι → ℤ` such that `f (b' i) = a i • b' i`. @@ -57,10 +60,10 @@ begin -- Now we map everything through the linear equiv `S ≃ₗ (ι → R)`, -- which maps `I` to `I' := Π i, a i ℤ`. - let I' : submodule R (ι → R) := submodule.pi set.univ (λ i, ideal.span ({a i} : set R)), + let I' : submodule R (ι → R) := submodule.pi set.univ (λ i, span ({a i} : set R)), have : submodule.map (b'.equiv_fun : S →ₗ[R] (ι → R)) (I.restrict_scalars R) = I', { ext x, - simp only [submodule.mem_map, submodule.mem_pi, ideal.mem_span_singleton, set.mem_univ, + simp only [submodule.mem_map, submodule.mem_pi, mem_span_singleton, set.mem_univ, submodule.restrict_scalars_mem, mem_I_iff, smul_eq_mul, forall_true_left, linear_equiv.coe_coe, basis.equiv_fun_apply], split, @@ -73,18 +76,18 @@ begin refine (submodule.quotient.equiv (I.restrict_scalars R) I' b'.equiv_fun this).trans _, any_goals { apply ring_hom.id }, any_goals { apply_instance }, classical, - let := submodule.quotient_pi (show Π i, submodule R R, from λ i, ideal.span ({a i} : set R)), + let := submodule.quotient_pi (show Π i, submodule R R, from λ i, span ({a i} : set R)), exact this end /-- Ideal quotients over a free finite extension of `ℤ` are isomorphic to a direct product of `zmod`. -/ -noncomputable def ideal.quotient_equiv_pi_zmod +noncomputable def quotient_equiv_pi_zmod (I : ideal S) (b : basis ι ℤ S) (hI : I ≠ ⊥) : (S ⧸ I) ≃+ Π i, (zmod (I.smith_coeffs b hI i).nat_abs) := let a := I.smith_coeffs b hI, e := I.quotient_equiv_pi_span b hI, - e' : (Π (i : ι), (ℤ ⧸ ideal.span ({a i} : set ℤ))) ≃+ Π (i : ι), zmod (a i).nat_abs := + e' : (Π (i : ι), (ℤ ⧸ span ({a i} : set ℤ))) ≃+ Π (i : ι), zmod (a i).nat_abs := add_equiv.Pi_congr_right (λ i, ↑(int.quotient_span_equiv_zmod (a i))) in (↑(e : (S ⧸ I) ≃ₗ[ℤ] _) : (S ⧸ I ≃+ _)).trans e' @@ -93,14 +96,14 @@ in (↑(e : (S ⧸ I) ≃ₗ[ℤ] _) : (S ⧸ I ≃+ _)).trans e' Can't be an instance because of the side condition `I ≠ ⊥`, and more importantly, because the choice of `fintype` instance is non-canonical. -/ -noncomputable def ideal.fintype_quotient_of_free_of_ne_bot [module.free ℤ S] [module.finite ℤ S] +noncomputable def fintype_quotient_of_free_of_ne_bot [module.free ℤ S] [module.finite ℤ S] (I : ideal S) (hI : I ≠ ⊥) : fintype (S ⧸ I) := let b := module.free.choose_basis ℤ S, a := I.smith_coeffs b hI, e := I.quotient_equiv_pi_zmod b hI in by haveI : ∀ i, ne_zero (a i).nat_abs := - (λ i, ⟨int.nat_abs_ne_zero_of_ne_zero (ideal.smith_coeffs_ne_zero b I hI i)⟩); classical; + (λ i, ⟨int.nat_abs_ne_zero_of_ne_zero (smith_coeffs_ne_zero b I hI i)⟩); classical; exact fintype.of_equiv (Π i, zmod (a i).nat_abs) e.symm variables (F : Type*) [comm_ring F] [algebra F R] [algebra F S] [is_scalar_tower F R S] @@ -108,9 +111,10 @@ variables (F : Type*) [comm_ring F] [algebra F R] [algebra F S] [is_scalar_tower /-- Decompose `S⧸I` as a direct sum of cyclic `R`-modules (quotients by the ideals generated by Smith coefficients of `I`). -/ -noncomputable def ideal.quotient_equiv_direct_sum : - (S ⧸ I) ≃ₗ[F] ⨁ i, R ⧸ ideal.span ({I.smith_coeffs b hI i} : set R) := +noncomputable def quotient_equiv_direct_sum : + (S ⧸ I) ≃ₗ[F] ⨁ i, R ⧸ span ({I.smith_coeffs b hI i} : set R) := begin + haveI := fintype.of_finite ι, apply ((I.quotient_equiv_pi_span b _).restrict_scalars F).trans (direct_sum.linear_equiv_fun_on_fintype _ _ _).symm, exact linear_map.is_scalar_tower.compatible_smul @@ -118,13 +122,15 @@ begin -- even after `change linear_map.compatible_smul _ (Π i, R ⧸ span _) F R` end -lemma ideal.finrank_quotient_eq_sum [nontrivial F] - [∀ i, module.free F (R ⧸ ideal.span ({I.smith_coeffs b hI i} : set R))] - [∀ i, module.finite F (R ⧸ ideal.span ({I.smith_coeffs b hI i} : set R))] : +lemma finrank_quotient_eq_sum {ι} [fintype ι] (b : basis ι R S) [nontrivial F] + [∀ i, module.free F (R ⧸ span ({I.smith_coeffs b hI i} : set R))] + [∀ i, module.finite F (R ⧸ span ({I.smith_coeffs b hI i} : set R))] : finite_dimensional.finrank F (S ⧸ I) - = ∑ i, finite_dimensional.finrank F (R ⧸ ideal.span ({I.smith_coeffs b hI i} : set R)) := + = ∑ i, finite_dimensional.finrank F (R ⧸ span ({I.smith_coeffs b hI i} : set R)) := begin - rw [linear_equiv.finrank_eq $ ideal.quotient_equiv_direct_sum F b hI, + rw [linear_equiv.finrank_eq $ quotient_equiv_direct_sum F b hI, finite_dimensional.finrank_direct_sum] -- slow, and dot notation doesn't work end + +end ideal diff --git a/src/linear_algebra/free_module/norm.lean b/src/linear_algebra/free_module/norm.lean new file mode 100644 index 0000000000000..7f73836649fac --- /dev/null +++ b/src/linear_algebra/free_module/norm.lean @@ -0,0 +1,75 @@ +/- +Copyright (c) 2023 Junyan Xu. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Junyan Xu +-/ + +import linear_algebra.free_module.ideal_quotient +import ring_theory.norm + +/-! +# Norms on free modules over principal ideal domains +-/ + +open ideal polynomial + +open_locale big_operators polynomial + +variables {R S ι : Type*} [comm_ring R] [is_domain R] [is_principal_ideal_ring R] [comm_ring S] + [is_domain S] [algebra R S] + +section comm_ring + +variables (F : Type*) [comm_ring F] [algebra F R] [algebra F S] [is_scalar_tower F R S] + +/-- For a nonzero element `f` in an algebra `S` over a principal ideal domain `R` that is finite and +free as an `R`-module, the norm of `f` relative to `R` is associated to the product of the Smith +coefficients of the ideal generated by `f`. -/ +lemma associated_norm_prod_smith [fintype ι] (b : basis ι R S) {f : S} (hf : f ≠ 0) : + associated (algebra.norm R f) (∏ i, smith_coeffs b _ (span_singleton_eq_bot.not.2 hf) i) := +begin + have hI := span_singleton_eq_bot.not.2 hf, + let b' := ring_basis b (span {f}) hI, + classical, + rw [← matrix.det_diagonal, ← linear_map.det_to_lin b'], + let e := (b'.equiv ((span {f}).self_basis b hI) $ equiv.refl _).trans + ((linear_equiv.coord S S f hf).restrict_scalars R), + refine (linear_map.associated_det_of_eq_comp e _ _ _).symm, + dsimp only [e, linear_equiv.trans_apply], + simp_rw [← linear_equiv.coe_to_linear_map, ← linear_map.comp_apply, ← linear_map.ext_iff], + refine b'.ext (λ i, _), + simp_rw [linear_map.comp_apply, linear_equiv.coe_to_linear_map, matrix.to_lin_apply, + basis.repr_self, finsupp.single_eq_pi_single, matrix.diagonal_mul_vec_single, pi.single_apply, + ite_smul, zero_smul, finset.sum_ite_eq', mul_one, if_pos (finset.mem_univ _), b'.equiv_apply], + change _ = f * _, + rw [mul_comm, ← smul_eq_mul, linear_equiv.restrict_scalars_apply, linear_equiv.coord_apply_smul, + ideal.self_basis_def], + refl +end + +end comm_ring + +section field + +variables {F : Type*} [field F] [algebra F[X] S] [finite ι] + +instance (b : basis ι F[X] S) {I : ideal S} (hI : I ≠ ⊥) (i : ι) : + finite_dimensional F (F[X] ⧸ span ({I.smith_coeffs b hI i} : set F[X])) := +(adjoin_root.power_basis $ I.smith_coeffs_ne_zero b hI i).finite_dimensional + +/-- For a nonzero element `f` in a `F[X]`-module `S`, the dimension of $S/\langle f \rangle$ as an +`F`-vector space is the degree of the norm of `f` relative to `F[X]`. -/ +lemma finrank_quotient_span_eq_nat_degree_norm [algebra F S] [is_scalar_tower F F[X] S] + (b : basis ι F[X] S) {f : S} (hf : f ≠ 0) : + finite_dimensional.finrank F (S ⧸ span ({f} : set S)) = (algebra.norm F[X] f).nat_degree := +begin + haveI := fintype.of_finite ι, + have h := span_singleton_eq_bot.not.2 hf, + rw [nat_degree_eq_of_degree_eq (degree_eq_degree_of_associated $ associated_norm_prod_smith b hf), + nat_degree_prod _ _ (λ i _, smith_coeffs_ne_zero b _ h i), finrank_quotient_eq_sum F h b], + -- finrank_quotient_eq_sum slow + congr' with i, + exact (adjoin_root.power_basis $ smith_coeffs_ne_zero b _ h i).finrank +end + +end field From c471da714c044131b90c133701e51b877c246677 Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Sat, 10 Jun 2023 20:00:22 +0000 Subject: [PATCH 1169/1291] chore(number_theory/legendre_symbol/gauss_eisenstein_lemmas): move zmod.wilsons_lemma (#19172) Co-authored-by: Scott Morrison --- .../gauss_eisenstein_lemmas.lean | 52 ++----------------- src/number_theory/wilson.lean | 50 ++++++++++++++++-- 2 files changed, 50 insertions(+), 52 deletions(-) diff --git a/src/number_theory/legendre_symbol/gauss_eisenstein_lemmas.lean b/src/number_theory/legendre_symbol/gauss_eisenstein_lemmas.lean index ce813915fe53b..75380ab52d02e 100644 --- a/src/number_theory/legendre_symbol/gauss_eisenstein_lemmas.lean +++ b/src/number_theory/legendre_symbol/gauss_eisenstein_lemmas.lean @@ -8,59 +8,13 @@ import number_theory.legendre_symbol.quadratic_reciprocity /-! # Lemmas of Gauss and Eisenstein -This file contains code for the proof of the Lemmas of Gauss and Eisenstein -on the Legendre symbol. The main results are `zmod.gauss_lemma_aux` and -`zmod.eisenstein_lemma_aux`. +This file contains the Lemmas of Gauss and Eisenstein on the Legendre symbol. +The main results are `zmod.gauss_lemma` and `zmod.eisenstein_lemma`. -/ -open function finset nat finite_field zmod +open finset nat open_locale big_operators nat -namespace zmod - -section wilson - -variables (p : ℕ) [fact p.prime] - -/-- **Wilson's Lemma**: the product of `1`, ..., `p-1` is `-1` modulo `p`. -/ -@[simp] lemma wilsons_lemma : ((p - 1)! : zmod p) = -1 := -begin - refine - calc ((p - 1)! : zmod p) = (∏ x in Ico 1 (succ (p - 1)), x) : - by rw [← finset.prod_Ico_id_eq_factorial, prod_nat_cast] - ... = (∏ x : (zmod p)ˣ, x) : _ - ... = -1 : by simp_rw [← units.coe_hom_apply, - ← (units.coe_hom (zmod p)).map_prod, prod_univ_units_id_eq_neg_one, units.coe_hom_apply, - units.coe_neg, units.coe_one], - have hp : 0 < p := (fact.out p.prime).pos, - symmetry, - refine prod_bij (λ a _, (a : zmod p).val) _ _ _ _, - { intros a ha, - rw [mem_Ico, ← nat.succ_sub hp, nat.succ_sub_one], - split, - { apply nat.pos_of_ne_zero, rw ← @val_zero p, - assume h, apply units.ne_zero a (val_injective p h) }, - { exact val_lt _ } }, - { intros a ha, simp only [cast_id, nat_cast_val], }, - { intros _ _ _ _ h, rw units.ext_iff, exact val_injective p h }, - { intros b hb, - rw [mem_Ico, nat.succ_le_iff, ← succ_sub hp, succ_sub_one, pos_iff_ne_zero] at hb, - refine ⟨units.mk0 b _, finset.mem_univ _, _⟩, - { assume h, apply hb.1, apply_fun val at h, - simpa only [val_cast_of_lt hb.right, val_zero] using h }, - { simp only [val_cast_of_lt hb.right, units.coe_mk0], } } -end - -@[simp] lemma prod_Ico_one_prime : (∏ x in Ico 1 p, (x : zmod p)) = -1 := -begin - conv in (Ico 1 p) { rw [← succ_sub_one p, succ_sub (fact.out p.prime).pos] }, - rw [← prod_nat_cast, finset.prod_Ico_id_eq_factorial, wilsons_lemma] -end - -end wilson - -end zmod - section gauss_eisenstein namespace zmod diff --git a/src/number_theory/wilson.lean b/src/number_theory/wilson.lean index e3a4f7c3f3a45..b253a9dbb3c90 100644 --- a/src/number_theory/wilson.lean +++ b/src/number_theory/wilson.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 John Nicol. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: John Nicol -/ -import number_theory.legendre_symbol.gauss_eisenstein_lemmas +import field_theory.finite.basic /-! # Wilson's theorem. @@ -21,10 +21,52 @@ This could be generalized to similar results about finite abelian groups. ## TODO -* Move `wilsons_lemma` into this file, and give it a descriptive name. +* Give `wilsons_lemma` a descriptive name. -/ -open_locale nat +open finset nat finite_field zmod +open_locale big_operators nat + +namespace zmod + +variables (p : ℕ) [fact p.prime] + +/-- **Wilson's Lemma**: the product of `1`, ..., `p-1` is `-1` modulo `p`. -/ +@[simp] lemma wilsons_lemma : ((p - 1)! : zmod p) = -1 := +begin + refine + calc ((p - 1)! : zmod p) = (∏ x in Ico 1 (succ (p - 1)), x) : + by rw [← finset.prod_Ico_id_eq_factorial, prod_nat_cast] + ... = (∏ x : (zmod p)ˣ, x) : _ + ... = -1 : by simp_rw [← units.coe_hom_apply, + ← (units.coe_hom (zmod p)).map_prod, prod_univ_units_id_eq_neg_one, units.coe_hom_apply, + units.coe_neg, units.coe_one], + have hp : 0 < p := (fact.out p.prime).pos, + symmetry, + refine prod_bij (λ a _, (a : zmod p).val) _ _ _ _, + { intros a ha, + rw [mem_Ico, ← nat.succ_sub hp, nat.succ_sub_one], + split, + { apply nat.pos_of_ne_zero, rw ← @val_zero p, + assume h, apply units.ne_zero a (val_injective p h) }, + { exact val_lt _ } }, + { intros a ha, simp only [cast_id, nat_cast_val], }, + { intros _ _ _ _ h, rw units.ext_iff, exact val_injective p h }, + { intros b hb, + rw [mem_Ico, nat.succ_le_iff, ← succ_sub hp, succ_sub_one, pos_iff_ne_zero] at hb, + refine ⟨units.mk0 b _, finset.mem_univ _, _⟩, + { assume h, apply hb.1, apply_fun val at h, + simpa only [val_cast_of_lt hb.right, val_zero] using h }, + { simp only [val_cast_of_lt hb.right, units.coe_mk0], } } +end + +@[simp] lemma prod_Ico_one_prime : (∏ x in Ico 1 p, (x : zmod p)) = -1 := +begin + conv in (Ico 1 p) { rw [← succ_sub_one p, succ_sub (fact.out p.prime).pos] }, + rw [← prod_nat_cast, finset.prod_Ico_id_eq_factorial, wilsons_lemma] +end + +end zmod namespace nat variable {n : ℕ} @@ -53,3 +95,5 @@ begin end end nat + +assert_not_exists legendre_sym.quadratic_reciprocity From 660b3a2db3522fa0db036e569dc995a615c4c848 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Sun, 11 Jun 2023 07:04:48 +0000 Subject: [PATCH 1170/1291] chore(*): add mathlib4 synchronization comments (#19174) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.lie.cartan_matrix` * `algebra.module.zlattice` * `analysis.complex.polynomial` * `category_theory.monoidal.limits` * `category_theory.preadditive.schur` * `category_theory.sites.compatible_sheafification` * `data.polynomial.unit_trinomial` * `field_theory.is_alg_closed.classification` * `field_theory.is_alg_closed.spectrum` * `linear_algebra.eigenspace.is_alg_closed` * `linear_algebra.matrix.charpoly.eigs` * `measure_theory.measure.portmanteau` --- src/algebra/lie/cartan_matrix.lean | 3 +++ src/algebra/module/zlattice.lean | 3 +++ src/analysis/complex/polynomial.lean | 3 +++ src/category_theory/monoidal/limits.lean | 3 +++ src/category_theory/preadditive/schur.lean | 3 +++ src/category_theory/sites/compatible_sheafification.lean | 3 +++ src/data/polynomial/unit_trinomial.lean | 3 +++ src/field_theory/is_alg_closed/classification.lean | 3 +++ src/field_theory/is_alg_closed/spectrum.lean | 3 +++ src/linear_algebra/eigenspace/is_alg_closed.lean | 3 +++ src/linear_algebra/matrix/charpoly/eigs.lean | 3 +++ src/measure_theory/measure/portmanteau.lean | 3 +++ 12 files changed, 36 insertions(+) diff --git a/src/algebra/lie/cartan_matrix.lean b/src/algebra/lie/cartan_matrix.lean index d3390999a509d..1e175daa8002c 100644 --- a/src/algebra/lie/cartan_matrix.lean +++ b/src/algebra/lie/cartan_matrix.lean @@ -10,6 +10,9 @@ import data.matrix.notation /-! # Lie algebras from Cartan matrices +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Split semi-simple Lie algebras are uniquely determined by their Cartan matrix. Indeed, if `A` is an `l × l` Cartan matrix, the corresponding Lie algebra may be obtained as the Lie algebra on `3l` generators: $H_1, H_2, \ldots H_l, E_1, E_2, \ldots, E_l, F_1, F_2, \ldots, F_l$ diff --git a/src/algebra/module/zlattice.lean b/src/algebra/module/zlattice.lean index 2d33d30ca3ea2..cdea3c81502dc 100644 --- a/src/algebra/module/zlattice.lean +++ b/src/algebra/module/zlattice.lean @@ -8,6 +8,9 @@ import measure_theory.group.fundamental_domain /-! # ℤ-lattices +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Let `E` be a finite dimensional vector space over a `normed_linear_ordered_field` `K` with a solid norm and that is also a `floor_ring`, e.g. `ℚ` or `ℝ`. A (full) ℤ-lattice `L` of `E` is a discrete subgroup of `E` such that `L` spans `E` over `K`. diff --git a/src/analysis/complex/polynomial.lean b/src/analysis/complex/polynomial.lean index d3a4710d9c922..db5074f4ea358 100644 --- a/src/analysis/complex/polynomial.lean +++ b/src/analysis/complex/polynomial.lean @@ -9,6 +9,9 @@ import field_theory.is_alg_closed.basic /-! # The fundamental theorem of algebra +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves that every nonconstant complex polynomial has a root using Liouville's theorem. As a consequence, the complex numbers are algebraically closed. diff --git a/src/category_theory/monoidal/limits.lean b/src/category_theory/monoidal/limits.lean index b8d72308365e4..930b3912659ab 100644 --- a/src/category_theory/monoidal/limits.lean +++ b/src/category_theory/monoidal/limits.lean @@ -10,6 +10,9 @@ import category_theory.limits.has_limits /-! # `lim : (J ⥤ C) ⥤ C` is lax monoidal when `C` is a monoidal category. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + When `C` is a monoidal category, the functorial association `F ↦ limit F` is lax monoidal, i.e. there are morphisms * `lim_lax.ε : (𝟙_ C) → limit (𝟙_ (J ⥤ C))` diff --git a/src/category_theory/preadditive/schur.lean b/src/category_theory/preadditive/schur.lean index f2a2125b233c6..4e46403e03014 100644 --- a/src/category_theory/preadditive/schur.lean +++ b/src/category_theory/preadditive/schur.lean @@ -11,6 +11,9 @@ import field_theory.is_alg_closed.spectrum /-! # Schur's lemma + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. We first prove the part of Schur's Lemma that holds in any preadditive category with kernels, that any nonzero morphism between simple objects is an isomorphism. diff --git a/src/category_theory/sites/compatible_sheafification.lean b/src/category_theory/sites/compatible_sheafification.lean index d9b682b87047a..23051dd93fe7d 100644 --- a/src/category_theory/sites/compatible_sheafification.lean +++ b/src/category_theory/sites/compatible_sheafification.lean @@ -7,6 +7,9 @@ import category_theory.sites.compatible_plus import category_theory.sites.sheafification /-! +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we prove that sheafification is compatible with functors which preserve the correct limits and colimits. diff --git a/src/data/polynomial/unit_trinomial.lean b/src/data/polynomial/unit_trinomial.lean index 08a659408f6c1..0a3957e21121e 100644 --- a/src/data/polynomial/unit_trinomial.lean +++ b/src/data/polynomial/unit_trinomial.lean @@ -10,6 +10,9 @@ import data.polynomial.mirror /-! # Unit Trinomials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines irreducible trinomials and proves an irreducibility criterion. ## Main definitions diff --git a/src/field_theory/is_alg_closed/classification.lean b/src/field_theory/is_alg_closed/classification.lean index f1efca9ea7eec..169081196a4db 100644 --- a/src/field_theory/is_alg_closed/classification.lean +++ b/src/field_theory/is_alg_closed/classification.lean @@ -11,6 +11,9 @@ import data.zmod.algebra /-! # Classification of Algebraically closed fields +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains results related to classifying algebraically closed fields. ## Main statements diff --git a/src/field_theory/is_alg_closed/spectrum.lean b/src/field_theory/is_alg_closed/spectrum.lean index f63674be31c92..8747b3ccb009f 100644 --- a/src/field_theory/is_alg_closed/spectrum.lean +++ b/src/field_theory/is_alg_closed/spectrum.lean @@ -9,6 +9,9 @@ import field_theory.is_alg_closed.basic /-! # Spectrum mapping theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file develops proves the spectral mapping theorem for polynomials over algebraically closed fields. In particular, if `a` is an element of an `𝕜`-algebra `A` where `𝕜` is a field, and `p : 𝕜[X]` is a polynomial, then the spectrum of `polynomial.aeval a p` contains the image of the diff --git a/src/linear_algebra/eigenspace/is_alg_closed.lean b/src/linear_algebra/eigenspace/is_alg_closed.lean index d8fb10800d9b6..1b896aa946d82 100644 --- a/src/linear_algebra/eigenspace/is_alg_closed.lean +++ b/src/linear_algebra/eigenspace/is_alg_closed.lean @@ -10,6 +10,9 @@ import field_theory.is_alg_closed.spectrum /-! # Eigenvectors and eigenvalues over algebraically closed fields. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + * Every linear operator on a vector space over an algebraically closed field has an eigenvalue. * The generalized eigenvectors span the entire vector space. diff --git a/src/linear_algebra/matrix/charpoly/eigs.lean b/src/linear_algebra/matrix/charpoly/eigs.lean index a585f16e77072..684f1bb8e94be 100644 --- a/src/linear_algebra/matrix/charpoly/eigs.lean +++ b/src/linear_algebra/matrix/charpoly/eigs.lean @@ -10,6 +10,9 @@ import field_theory.is_alg_closed.basic /-! # Eigenvalues are characteristic polynomial roots. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In fields we show that: * `matrix.det_eq_prod_roots_charpoly_of_splits`: the determinant (in the field of the matrix) diff --git a/src/measure_theory/measure/portmanteau.lean b/src/measure_theory/measure/portmanteau.lean index 24f012aeba70e..12f92076938ed 100644 --- a/src/measure_theory/measure/portmanteau.lean +++ b/src/measure_theory/measure/portmanteau.lean @@ -9,6 +9,9 @@ import measure_theory.measure.lebesgue.basic /-! # Characterizations of weak convergence of finite measures and probability measures +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file will provide portmanteau characterizations of the weak convergence of finite measures and of probability measures, i.e., the standard characterizations of convergence in distribution. From ccdbfb6e5614667af5aa3ab2d50885e0ef44a46f Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sun, 11 Jun 2023 12:27:24 +0000 Subject: [PATCH 1171/1291] feat(measure_theory/integral/set_integral): First moment method (#18731) Integrable functions are smaller/larger than their mean on a set of positive measure. We prove it for the Bochner and Lebesgue integrals. --- src/analysis/mean_inequalities_pow.lean | 2 +- src/data/real/ennreal.lean | 11 +- src/measure_theory/function/l1_space.lean | 2 +- src/measure_theory/integral/average.lean | 213 +++++++++++++++++++++- 4 files changed, 220 insertions(+), 8 deletions(-) diff --git a/src/analysis/mean_inequalities_pow.lean b/src/analysis/mean_inequalities_pow.lean index f63af0fd6e1de..e639c518c2bff 100644 --- a/src/analysis/mean_inequalities_pow.lean +++ b/src/analysis/mean_inequalities_pow.lean @@ -297,7 +297,7 @@ begin { simp [← mul_assoc, ennreal.inv_mul_cancel two_ne_zero two_ne_top] }, { have A : p - 1 ≠ 0 := ne_of_gt (sub_pos.2 h'p), simp only [mul_rpow_of_nonneg _ _ (zero_le_one.trans hp), rpow_sub _ _ two_ne_zero two_ne_top, - div_eq_inv_mul, rpow_one, mul_one], + ennreal.div_eq_inv_mul, rpow_one, mul_one], ring } end diff --git a/src/data/real/ennreal.lean b/src/data/real/ennreal.lean index f54666b723bf6..00d67e43847e3 100644 --- a/src/data/real/ennreal.lean +++ b/src/data/real/ennreal.lean @@ -193,6 +193,12 @@ by rintro (h | h); simp [h]⟩ lemma to_real_eq_zero_iff (x : ℝ≥0∞) : x.to_real = 0 ↔ x = 0 ∨ x = ∞ := by simp [ennreal.to_real, to_nnreal_eq_zero_iff] +lemma to_nnreal_ne_zero : a.to_nnreal ≠ 0 ↔ a ≠ 0 ∧ a ≠ ∞ := +a.to_nnreal_eq_zero_iff.not.trans not_or_distrib + +lemma to_real_ne_zero : a.to_real ≠ 0 ↔ a ≠ 0 ∧ a ≠ ∞ := +a.to_real_eq_zero_iff.not.trans not_or_distrib + lemma to_nnreal_eq_one_iff (x : ℝ≥0∞) : x.to_nnreal = 1 ↔ x = 1 := begin refine ⟨λ h, _, congr_arg _⟩, @@ -204,6 +210,9 @@ end lemma to_real_eq_one_iff (x : ℝ≥0∞) : x.to_real = 1 ↔ x = 1 := by rw [ennreal.to_real, nnreal.coe_eq_one, ennreal.to_nnreal_eq_one_iff] +lemma to_nnreal_ne_one : a.to_nnreal ≠ 1 ↔ a ≠ 1 := a.to_nnreal_eq_one_iff.not +lemma to_real_ne_one : a.to_real ≠ 1 ↔ a ≠ 1 := a.to_real_eq_one_iff.not + @[simp] lemma coe_ne_top : (r : ℝ≥0∞) ≠ ∞ := with_top.coe_ne_top @[simp] lemma top_ne_coe : ∞ ≠ (r : ℝ≥0∞) := with_top.top_ne_coe @[simp] lemma of_real_ne_top {r : ℝ} : ennreal.of_real r ≠ ∞ := by simp [ennreal.of_real] @@ -996,7 +1005,7 @@ instance : div_inv_monoid ℝ≥0∞ := { inv := has_inv.inv, .. (infer_instance : monoid ℝ≥0∞) } -lemma div_eq_inv_mul : a / b = b⁻¹ * a := by rw [div_eq_mul_inv, mul_comm] +protected lemma div_eq_inv_mul : a / b = b⁻¹ * a := by rw [div_eq_mul_inv, mul_comm] @[simp] lemma inv_zero : (0 : ℝ≥0∞)⁻¹ = ∞ := show Inf {b : ℝ≥0∞ | 1 ≤ 0 * b} = ∞, by simp; refl diff --git a/src/measure_theory/function/l1_space.lean b/src/measure_theory/function/l1_space.lean index 30a1a57d68d80..0ca5e2f3fd606 100644 --- a/src/measure_theory/function/l1_space.lean +++ b/src/measure_theory/function/l1_space.lean @@ -456,7 +456,7 @@ begin rw [integrable, and_iff_right this, has_finite_integral_const_iff] end -lemma integrable_const [is_finite_measure μ] (c : β) : integrable (λ x : α, c) μ := +@[simp] 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∞} diff --git a/src/measure_theory/integral/average.lean b/src/measure_theory/integral/average.lean index efd336ea1c821..230993042e849 100644 --- a/src/measure_theory/integral/average.lean +++ b/src/measure_theory/integral/average.lean @@ -1,7 +1,7 @@ /- 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 +Authors: Yury G. Kudryashov, Yaël Dillies -/ import measure_theory.integral.set_integral @@ -19,24 +19,32 @@ measure, then the average of any function is equal to its integral. For the average on a set, we use `⨍ x in s, f x ∂μ` (notation for `⨍ x, f x ∂(μ.restrict s)`). For average w.r.t. the volume, one can omit `∂volume`. +We prove several version of the first moment method: An integrable function is below/above its +average on a set of positive measure. + ## Implementation notes The average is defined as an integral over `(μ univ)⁻¹ • μ` so that all theorems about Bochner integrals work for the average without modifications. For theorems that require integrability of a function, we provide a convenience lemma `measure_theory.integrable.to_average`. +## TODO + +Provide the first moment method for the Lebesgue integral as well. A draft is available on branch +`first_moment_lintegral`. + ## Tags integral, center mass, average value -/ -open measure_theory measure_theory.measure metric set filter topological_space function +open ennreal measure_theory measure_theory.measure metric set filter topological_space function open_locale topology big_operators ennreal convex variables {α E F : Type*} {m0 : measurable_space α} [normed_add_comm_group E] [normed_space ℝ E] [complete_space E] [normed_add_comm_group F] [normed_space ℝ F] [complete_space F] - {μ : measure α} {s : set E} + {μ : measure α} {s t : set α} /-! ### Average value of a function w.r.t. a measure @@ -49,8 +57,8 @@ integral. -/ namespace measure_theory - -variable (μ) +section normed_add_comm_group +variables (μ) {f g : α → E} include m0 /-- Average value of a function `f` w.r.t. a measure `μ`, notation: `⨍ x, f x ∂μ`. It is defined as @@ -106,6 +114,9 @@ variable {μ} lemma average_congr {f g : α → E} (h : f =ᵐ[μ] g) : ⨍ x, f x ∂μ = ⨍ x, g x ∂μ := by simp only [average_eq, integral_congr_ae h] +lemma set_average_congr_set_ae (h : s =ᵐ[μ] t) : ⨍ x in s, f x ∂μ = ⨍ x in t, f x ∂μ := +by simp only [set_average_eq, set_integral_congr_set_ae h, measure_congr h] + lemma average_add_measure [is_finite_measure μ] {ν : measure α} [is_finite_measure ν] {f : α → E} (hμ : integrable f μ) (hν : integrable f ν) : ⨍ x, f x ∂(μ + ν) = @@ -182,4 +193,196 @@ by simp only [set_average_eq, integral_const, measure.restrict_apply, measurable univ_inter, smul_smul, ← ennreal.to_real_inv, ← ennreal.to_real_mul, ennreal.inv_mul_cancel hs₀ hs, ennreal.one_to_real, one_smul] +@[simp] lemma integral_average (μ : measure α) [is_finite_measure μ] (f : α → E) : + ∫ x, ⨍ a, f a ∂μ ∂μ = ∫ x, f x ∂μ := +begin + unfreezingI { obtain rfl | hμ := eq_or_ne μ 0 }, + { simp only [integral_zero_measure] }, + { rw [integral_const, average_eq, + smul_inv_smul₀ (to_real_ne_zero.2 ⟨measure_univ_ne_zero.2 hμ, measure_ne_top _ _⟩)] } +end + +lemma set_integral_set_average (μ : measure α) [is_finite_measure μ] (f : α → E) (s : set α) : + ∫ x in s, ⨍ a in s, f a ∂μ ∂μ = ∫ x in s, f x ∂μ := +integral_average _ _ + +lemma integral_sub_average (μ : measure α) [is_finite_measure μ] (f : α → E) : + ∫ x, f x - ⨍ a, f a ∂μ ∂μ = 0 := +begin + by_cases hf : integrable f μ, + { rw [integral_sub hf (integrable_const _), integral_average, sub_self] }, + refine integral_undef (λ h, hf _), + convert h.add (integrable_const _), + exact (sub_add_cancel _ _).symm, +end + +lemma set_integral_sub_set_average (hs : μ s ≠ ∞) (f : α → E) : + ∫ x in s, f x - ⨍ a in s, f a ∂μ ∂μ = 0 := +by haveI haveI : fact (μ s < ∞) := ⟨lt_top_iff_ne_top.2 hs⟩; exact integral_sub_average _ _ + +lemma integral_average_sub [is_finite_measure μ] (hf : integrable f μ) : + ∫ x, ⨍ a, f a ∂μ - f x ∂μ = 0 := +by rw [integral_sub (integrable_const _) hf, integral_average, sub_self] + +lemma set_integral_set_average_sub (hs : μ s ≠ ∞) (hf : integrable_on f s μ) : + ∫ x in s, ⨍ a in s, f a ∂μ - f x ∂μ = 0 := +by haveI haveI : fact (μ s < ∞) := ⟨lt_top_iff_ne_top.2 hs⟩; exact integral_average_sub hf + +end normed_add_comm_group + +lemma of_real_average {f : α → ℝ} (hf : integrable f μ) (hf₀ : 0 ≤ᵐ[μ] f) : + ennreal.of_real (⨍ x, f x ∂μ) = (∫⁻ x, ennreal.of_real (f x) ∂μ) / μ univ := +begin + obtain rfl | hμ := eq_or_ne μ 0, + { simp }, + { rw [average_eq, smul_eq_mul, ←to_real_inv, of_real_mul (to_real_nonneg), + of_real_to_real (inv_ne_top.2 $ measure_univ_ne_zero.2 hμ), + of_real_integral_eq_lintegral_of_real hf hf₀, ennreal.div_eq_inv_mul] } +end + +lemma of_real_set_average {f : α → ℝ} (hf : integrable_on f s μ) + (hf₀ : 0 ≤ᵐ[μ.restrict s] f) : + ennreal.of_real (⨍ x in s, f x ∂μ) = (∫⁻ x in s, ennreal.of_real (f x) ∂μ) / μ s := +by simpa using of_real_average hf hf₀ + +lemma average_to_real {f : α → ℝ≥0∞} (hf : ae_measurable f μ) (hf' : ∀ᵐ x ∂μ, f x ≠ ∞) : + ⨍ x, (f x).to_real ∂μ = (∫⁻ x, f x ∂μ / μ univ).to_real := +begin + obtain rfl | hμ := eq_or_ne μ 0, + { simp }, + { rw [average_eq, smul_eq_mul, to_real_div, + ←integral_to_real hf (hf'.mono $ λ _, lt_top_iff_ne_top.2), div_eq_inv_mul] } +end + +lemma set_average_to_real {f : α → ℝ≥0∞} (hf : ae_measurable f (μ.restrict s)) + (hf' : ∀ᵐ x ∂(μ.restrict s), f x ≠ ∞) : + ⨍ x in s, (f x).to_real ∂μ = (∫⁻ x in s, f x ∂μ / μ s).to_real := +by simpa using average_to_real hf hf' + +/-! ### First moment method -/ + +section first_moment +variables {N : set α} {f : α → ℝ} + +/-- **First moment method**. An integrable function is smaller than its mean on a set of positive +measure. -/ +lemma measure_le_set_average_pos (hμ : μ s ≠ 0) (hμ₁ : μ s ≠ ∞) (hf : integrable_on f s μ) : + 0 < μ {x ∈ s | f x ≤ ⨍ a in s, f a ∂μ} := +begin + refine pos_iff_ne_zero.2 (λ H, _), + replace H : (μ.restrict s) {x | f x ≤ ⨍ a in s, f a ∂μ} = 0, + { rwa [restrict_apply₀, inter_comm], + exact ae_strongly_measurable.null_measurable_set_le hf.1 ae_strongly_measurable_const }, + haveI : is_finite_measure (μ.restrict s) := + ⟨by simpa only [measure.restrict_apply, measurable_set.univ, univ_inter] using hμ₁.lt_top⟩, + refine (integral_sub_average (μ.restrict s) f).not_gt _, + refine (set_integral_pos_iff_support_of_nonneg_ae _ _).2 _, + { refine eq_bot_mono (measure_mono (λ x hx, _)) H, + simp only [pi.zero_apply, sub_nonneg, mem_compl_iff, mem_set_of_eq, not_le] at hx, + exact hx.le }, + { exact hf.sub (integrable_on_const.2 $ or.inr $ lt_top_iff_ne_top.2 hμ₁) }, + { rwa [pos_iff_ne_zero, inter_comm, ←diff_compl, ←diff_inter_self_eq_diff, measure_diff_null], + refine eq_bot_mono (measure_mono _) (measure_inter_eq_zero_of_restrict H), + exact inter_subset_inter_left _ (λ a ha, (sub_eq_zero.1 $ of_not_not ha).le) } +end + +/-- **First moment method**. An integrable function is greater than its mean on a set of positive +measure. -/ +lemma measure_set_average_le_pos (hμ : μ s ≠ 0) (hμ₁ : μ s ≠ ∞) (hf : integrable_on f s μ) : + 0 < μ {x ∈ s | ⨍ a in s, f a ∂μ ≤ f x} := +by simpa [integral_neg, neg_div] using measure_le_set_average_pos hμ hμ₁ hf.neg + +/-- **First moment method**. The minimum of an integrable function is smaller than its mean. -/ +lemma exists_le_set_average (hμ : μ s ≠ 0) (hμ₁ : μ s ≠ ∞) (hf : integrable_on f s μ) : + ∃ x ∈ s, f x ≤ ⨍ a in s, f a ∂μ := +let ⟨x, hx, h⟩ := nonempty_of_measure_ne_zero (measure_le_set_average_pos hμ hμ₁ hf).ne' + in ⟨x, hx, h⟩ + +/-- **First moment method**. The maximum of an integrable function is greater than its mean. -/ +lemma exists_set_average_le (hμ : μ s ≠ 0) (hμ₁ : μ s ≠ ∞) (hf : integrable_on f s μ) : + ∃ x ∈ s, ⨍ a in s, f a ∂μ ≤ f x := +let ⟨x, hx, h⟩ := nonempty_of_measure_ne_zero (measure_set_average_le_pos hμ hμ₁ hf).ne' + in ⟨x, hx, h⟩ + +section finite_measure +variables [is_finite_measure μ] + +/-- **First moment method**. An integrable function is smaller than its mean on a set of positive +measure. -/ +lemma measure_le_average_pos (hμ : μ ≠ 0) (hf : integrable f μ) : 0 < μ {x | f x ≤ ⨍ a, f a ∂μ} := +by simpa using measure_le_set_average_pos (measure.measure_univ_ne_zero.2 hμ) (measure_ne_top _ _) + hf.integrable_on + +/-- **First moment method**. An integrable function is greater than its mean on a set of positive +measure. -/ +lemma measure_average_le_pos (hμ : μ ≠ 0) (hf : integrable f μ) : 0 < μ {x | ⨍ a, f a ∂μ ≤ f x} := +by simpa using measure_set_average_le_pos (measure.measure_univ_ne_zero.2 hμ) (measure_ne_top _ _) + hf.integrable_on + +/-- **First moment method**. The minimum of an integrable function is smaller than its mean. -/ +lemma exists_le_average (hμ : μ ≠ 0) (hf : integrable f μ) : ∃ x, f x ≤ ⨍ a, f a ∂μ := +let ⟨x, hx⟩ := nonempty_of_measure_ne_zero (measure_le_average_pos hμ hf).ne' in ⟨x, hx⟩ + +/-- **First moment method**. The maximum of an integrable function is greater than its mean. -/ +lemma exists_average_le (hμ : μ ≠ 0) (hf : integrable f μ) : ∃ x, ⨍ a, f a ∂μ ≤ f x := +let ⟨x, hx⟩ := nonempty_of_measure_ne_zero (measure_average_le_pos hμ hf).ne' in ⟨x, hx⟩ + +/-- **First moment method**. The minimum of an integrable function is smaller than its mean, while +avoiding a null set. -/ +lemma exists_not_mem_null_le_average (hμ : μ ≠ 0) (hf : integrable f μ) (hN : μ N = 0) : + ∃ x ∉ N, f x ≤ ⨍ a, f a ∂μ := +begin + have := measure_le_average_pos hμ hf, + rw ←measure_diff_null hN at this, + obtain ⟨x, hx, hxN⟩ := nonempty_of_measure_ne_zero this.ne', + exact ⟨x, hxN, hx⟩, +end + +/-- **First moment method**. The maximum of an integrable function is greater than its mean, while +avoiding a null set. -/ +lemma exists_not_mem_null_average_le (hμ : μ ≠ 0) (hf : integrable f μ) (hN : μ N = 0) : + ∃ x ∉ N, ⨍ a, f a ∂μ ≤ f x := +by simpa [integral_neg, neg_div] using exists_not_mem_null_le_average hμ hf.neg hN + +end finite_measure + +section probability_measure +variables [is_probability_measure μ] + +/-- **First moment method**. An integrable function is smaller than its integral on a set of +positive measure. -/ +lemma measure_le_integral_pos (hf : integrable f μ) : 0 < μ {x | f x ≤ ∫ a, f a ∂μ} := +by simpa only [average_eq_integral] + using measure_le_average_pos (is_probability_measure.ne_zero μ) hf + +/-- **First moment method**. An integrable function is greater than its integral on a set of +positive measure. -/ +lemma measure_integral_le_pos (hf : integrable f μ) : 0 < μ {x | ∫ a, f a ∂μ ≤ f x} := +by simpa only [average_eq_integral] + using measure_average_le_pos (is_probability_measure.ne_zero μ) hf + +/-- **First moment method**. The minimum of an integrable function is smaller than its integral. -/ +lemma exists_le_integral (hf : integrable f μ) : ∃ x, f x ≤ ∫ a, f a ∂μ := +by simpa only [average_eq_integral] using exists_le_average (is_probability_measure.ne_zero μ) hf + +/-- **First moment method**. The maximum of an integrable function is greater than its integral. -/ +lemma exists_integral_le (hf : integrable f μ) : ∃ x, ∫ a, f a ∂μ ≤ f x := +by simpa only [average_eq_integral] using exists_average_le (is_probability_measure.ne_zero μ) hf + +/-- **First moment method**. The minimum of an integrable function is smaller than its integral, +while avoiding a null set. -/ +lemma exists_not_mem_null_le_integral (hf : integrable f μ) (hN : μ N = 0) : + ∃ x ∉ N, f x ≤ ∫ a, f a ∂μ := +by simpa only [average_eq_integral] + using exists_not_mem_null_le_average (is_probability_measure.ne_zero μ) hf hN + +/-- **First moment method**. The maximum of an integrable function is greater than its integral, +while avoiding a null set. -/ +lemma exists_not_mem_null_integral_le (hf : integrable f μ) (hN : μ N = 0) : + ∃ x ∉ N, ∫ a, f a ∂μ ≤ f x := +by simpa only [average_eq_integral] + using exists_not_mem_null_average_le (is_probability_measure.ne_zero μ) hf hN + +end probability_measure +end first_moment end measure_theory From 7fdd4f3746cb059edfdb5d52cba98f66fce418c0 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sun, 11 Jun 2023 22:44:02 +0000 Subject: [PATCH 1172/1291] refactor(data/nat/nth): redefine, review API (#18760) Redefine `nat.nth` in terms of already available definitions, review API, generalize some lemmas. Also fix some typos in `data/set/intervals/monotone`. --- src/data/nat/nth.lean | 567 ++++++++++++-------------- src/data/set/finite.lean | 8 + src/number_theory/prime_counting.lean | 4 +- 3 files changed, 278 insertions(+), 301 deletions(-) diff --git a/src/data/nat/nth.lean b/src/data/nat/nth.lean index 18e2db160b439..c7802520e6ce0 100644 --- a/src/data/nat/nth.lean +++ b/src/data/nat/nth.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies, Vladimir Goryachev, Kyle Miller, Scott Morrison, Eric Rodriguez -/ import data.nat.count +import data.set.intervals.monotone import order.order_iso_nat /-! @@ -14,14 +15,14 @@ and provides lemmas that deal with this function and its connection to `nat.coun ## Main definitions -* `nth p n`: The `n`-th natural `k` (zero-indexed) such that `p k`. If there is no - such natural (that is, `p` is true for at most `n` naturals), then `nth p n = 0`. +* `nat.nth p n`: The `n`-th natural `k` (zero-indexed) such that `p k`. If there is no + such natural (that is, `p` is true for at most `n` naturals), then `nat.nth p n = 0`. ## Main results * `nat.nth_set_card`: For a fintely-often true `p`, gives the cardinality of the set of numbers satisfying `p` above particular values of `nth p` -* `nat.count_nth_gc`: Establishes a Galois connection between `nth p` and `count p`. +* `nat.count_nth_gc`: Establishes a Galois connection between `nat.nth p` and `nat.count p`. * `nat.nth_eq_order_iso_of_nat`: For an infinitely-ofter true predicate, `nth` agrees with the order-isomorphism of the subtype to the natural numbers. @@ -40,367 +41,335 @@ variable (p : ℕ → Prop) /-- Find the `n`-th natural number satisfying `p` (indexed from `0`, so `nth p 0` is the first natural number satisfying `p`), or `0` if there is no such number. See also `subtype.order_iso_of_nat` for the order isomorphism with ℕ when `p` is infinitely often true. -/ -noncomputable def nth : ℕ → ℕ -| n := Inf { i : ℕ | p i ∧ ∀ k < n, nth k < i } +noncomputable def nth (p : ℕ → Prop) (n : ℕ) : ℕ := +by classical; exact + if h : set.finite (set_of p) then (h.to_finset.sort (≤)).nthd n 0 + else @nat.subtype.order_iso_of_nat (set_of p) (set.infinite.to_subtype h) n -lemma nth_zero : nth p 0 = Inf { i : ℕ | p i } := by { rw nth, simp } +variable {p} -@[simp] lemma nth_zero_of_zero (h : p 0) : nth p 0 = 0 := -by simp [nth_zero, h] +/-! +### Lemmas about `nat.nth` on a finite set +-/ -lemma nth_zero_of_exists [decidable_pred p] (h : ∃ n, p n) : nth p 0 = nat.find h := -by { rw [nth_zero], convert nat.Inf_def h } +theorem nth_of_card_le (hf : (set_of p).finite) {n : ℕ} (hn : hf.to_finset.card ≤ n) : + nth p n = 0 := +by { rw [nth, dif_pos hf, list.nthd_eq_default], rwa [finset.length_sort] } -lemma nth_set_card_aux {n : ℕ} (hp : (set_of p).finite) - (hp' : {i : ℕ | p i ∧ ∀ t < n, nth p t < i}.finite) (hle : n ≤ hp.to_finset.card) : - hp'.to_finset.card = hp.to_finset.card - n := -begin - unfreezingI { induction n with k hk }, - { congr, - simp only [is_empty.forall_iff, nat.not_lt_zero, forall_const, and_true] }, - have hp'': {i : ℕ | p i ∧ ∀ t, t < k → nth p t < i}.finite, - { refine hp.subset (λ x hx, _), - rw set.mem_set_of_eq at hx, - exact hx.left }, - have hle' := nat.sub_pos_of_lt hle, - specialize hk hp'' (k.le_succ.trans hle), - rw [nat.sub_succ', ←hk], - convert_to (finset.erase hp''.to_finset (nth p k)).card = _, - { congr, - ext a, - simp only [set.finite.mem_to_finset, ne.def, set.mem_set_of_eq, finset.mem_erase], - refine ⟨λ ⟨hp, hlt⟩, - ⟨(hlt _ (lt_add_one k)).ne', ⟨hp, λ n hn, hlt n (hn.trans_le k.le_succ)⟩⟩, _⟩, - rintro ⟨hak : _ ≠ _, hp, hlt⟩, - refine ⟨hp, λ n hn, _⟩, - rw lt_succ_iff at hn, - obtain hnk | rfl := hn.lt_or_eq, - { exact hlt _ hnk }, - { refine lt_of_le_of_ne _ (ne.symm hak), - rw nth, - apply nat.Inf_le, - simpa [hp] using hlt } }, - apply finset.card_erase_of_mem, - rw [nth, set.finite.mem_to_finset], - apply Inf_mem, - rwa [←hp''.to_finset_nonempty, ←finset.card_pos, hk], -end +theorem nth_eq_nthd_sort (h : (set_of p).finite) (n : ℕ) : + nth p n = (h.to_finset.sort (≤)).nthd n 0 := +dif_pos h -lemma nth_set_card {n : ℕ} (hp : (set_of p).finite) - (hp' : {i : ℕ | p i ∧ ∀ k < n, nth p k < i}.finite) : - hp'.to_finset.card = hp.to_finset.card - n := -begin - obtain hn | hn := le_or_lt n hp.to_finset.card, - { exact nth_set_card_aux p hp _ hn }, - rw nat.sub_eq_zero_of_le hn.le, - simp only [finset.card_eq_zero, set.finite.to_finset_eq_empty, ←set.subset_empty_iff], - convert_to _ ⊆ {i : ℕ | p i ∧ ∀ (k : ℕ), k < hp.to_finset.card → nth p k < i}, - { symmetry, - rw [←set.finite.to_finset_eq_empty, ←finset.card_eq_zero, - ←nat.sub_self hp.to_finset.card], - { apply nth_set_card_aux p hp _ le_rfl }, - { exact hp.subset (λ x hx, hx.1) } }, - exact λ x hx, ⟨hx.1, λ k hk, hx.2 _ (hk.trans hn)⟩, -end +theorem nth_eq_order_emb_of_fin (hf : (set_of p).finite) {n : ℕ} (hn : n < hf.to_finset.card) : + nth p n = hf.to_finset.order_emb_of_fin rfl ⟨n, hn⟩ := +by { rw [nth_eq_nthd_sort hf, finset.order_emb_of_fin_apply, list.nthd_eq_nth_le], refl } -lemma nth_set_nonempty_of_lt_card {n : ℕ} (hp : (set_of p).finite) (hlt: n < hp.to_finset.card) : - {i : ℕ | p i ∧ ∀ k < n, nth p k < i}.nonempty := +theorem nth_strict_mono_on (hf : (set_of p).finite) : + strict_mono_on (nth p) (set.Iio hf.to_finset.card) := begin - have hp': {i : ℕ | p i ∧ ∀ (k : ℕ), k < n → nth p k < i}.finite, - { exact hp.subset (λ x hx, hx.1) }, - rw [←hp'.to_finset_nonempty, ←finset.card_pos, nth_set_card p hp], - exact nat.sub_pos_of_lt hlt, + rintro m (hm : m < _) n (hn : n < _) h, + simp only [nth_eq_order_emb_of_fin, *], + exact order_embedding.strict_mono _ h end -lemma nth_mem_of_lt_card_finite_aux (n : ℕ) (hp : (set_of p).finite) (hlt : n < hp.to_finset.card) : - nth p n ∈ {i : ℕ | p i ∧ ∀ k < n, nth p k < i} := +theorem nth_lt_nth_of_lt_card (hf : (set_of p).finite) {m n : ℕ} (h : m < n) + (hn : n < hf.to_finset.card) : nth p m < nth p n := +nth_strict_mono_on hf (h.trans hn) hn h + +theorem nth_le_nth_of_lt_card (hf : (set_of p).finite) {m n : ℕ} (h : m ≤ n) + (hn : n < hf.to_finset.card) : nth p m ≤ nth p n := +(nth_strict_mono_on hf).monotone_on (h.trans_lt hn) hn h + +theorem lt_of_nth_lt_nth_of_lt_card (hf : (set_of p).finite) {m n : ℕ} (h : nth p m < nth p n) + (hm : m < hf.to_finset.card) : m < n := +not_le.1 $ λ hle, h.not_le $ nth_le_nth_of_lt_card hf hle hm + +theorem le_of_nth_le_nth_of_lt_card (hf : (set_of p).finite) {m n : ℕ} (h : nth p m ≤ nth p n) + (hm : m < hf.to_finset.card) : m ≤ n := +not_lt.1 $ λ hlt, h.not_lt $ nth_lt_nth_of_lt_card hf hlt hm + +theorem nth_inj_on (hf : (set_of p).finite) : (set.Iio hf.to_finset.card).inj_on (nth p) := +(nth_strict_mono_on hf).inj_on + +theorem range_nth_of_finite (hf : (set_of p).finite) : set.range (nth p) = insert 0 (set_of p) := +by simpa only [← nth_eq_nthd_sort hf, mem_sort, set.finite.mem_to_finset] + using set.range_list_nthd (hf.to_finset.sort (≤)) 0 + +@[simp] theorem image_nth_Iio_card (hf : (set_of p).finite) : + nth p '' set.Iio hf.to_finset.card = set_of p := +calc nth p '' set.Iio hf.to_finset.card = set.range (hf.to_finset.order_emb_of_fin rfl) : + by ext x; simp only [set.mem_image, set.mem_range, fin.exists_iff, + ← nth_eq_order_emb_of_fin hf, set.mem_Iio, exists_prop] +... = set_of p : by rw [range_order_emb_of_fin, set.finite.coe_to_finset] + +lemma nth_mem_of_lt_card {n : ℕ} (hf : (set_of p).finite) (hlt : n < hf.to_finset.card) : + p (nth p n) := +(image_nth_Iio_card hf).subset $ set.mem_image_of_mem _ hlt + +theorem exists_lt_card_finite_nth_eq (hf : (set_of p).finite) {x} (h : p x) : + ∃ n, n < hf.to_finset.card ∧ nth p n = x := +by rwa [← @set.mem_set_of_eq _ _ p, ← image_nth_Iio_card hf] at h + +/-! +### Lemmas about `nat.nth` on an infinite set +-/ + +/-- When `s` is an infinite set, `nth` agrees with `nat.subtype.order_iso_of_nat`. -/ +theorem nth_apply_eq_order_iso_of_nat (hf : (set_of p).infinite) (n : ℕ) : + nth p n = @nat.subtype.order_iso_of_nat (set_of p) hf.to_subtype n := +by rw [nth, dif_neg hf] + +/-- When `s` is an infinite set, `nth` agrees with `nat.subtype.order_iso_of_nat`. -/ +theorem nth_eq_order_iso_of_nat (hf : (set_of p).infinite) : + nth p = coe ∘ @nat.subtype.order_iso_of_nat (set_of p) hf.to_subtype := +funext $ nth_apply_eq_order_iso_of_nat hf + +lemma nth_strict_mono (hf : (set_of p).infinite) : strict_mono (nth p) := begin - rw nth, - apply Inf_mem, - exact nth_set_nonempty_of_lt_card _ _ hlt, + rw [nth_eq_order_iso_of_nat hf], + exact (subtype.strict_mono_coe _).comp (order_iso.strict_mono _) end -lemma nth_mem_of_lt_card_finite {n : ℕ} (hp : (set_of p).finite) (hlt : n < hp.to_finset.card) : - p (nth p n) := (nth_mem_of_lt_card_finite_aux p n hp hlt).1 +lemma nth_injective (hf : (set_of p).infinite) : function.injective (nth p) := +(nth_strict_mono hf).injective + +lemma nth_monotone (hf : (set_of p).infinite) : monotone (nth p) := +(nth_strict_mono hf).monotone -lemma nth_strict_mono_of_finite {m n : ℕ} (hp : (set_of p).finite) - (hlt : n < hp.to_finset.card) (hmn : m < n) : nth p m < nth p n := -(nth_mem_of_lt_card_finite_aux p _ hp hlt).2 _ hmn +lemma nth_lt_nth (hf : (set_of p).infinite) {k n} : nth p k < nth p n ↔ k < n := +(nth_strict_mono hf).lt_iff_lt -lemma nth_mem_of_infinite_aux (hp : (set_of p).infinite) (n : ℕ) : - nth p n ∈ { i : ℕ | p i ∧ ∀ k < n, nth p k < i } := +lemma nth_le_nth (hf : (set_of p).infinite) {k n} : nth p k ≤ nth p n ↔ k ≤ n := +(nth_strict_mono hf).le_iff_le + +lemma range_nth_of_infinite (hf : (set_of p).infinite) : set.range (nth p) = set_of p := begin - rw nth, - apply Inf_mem, - let s : set ℕ := ⋃ (k < n), { i : ℕ | nth p k ≥ i }, - convert_to ((set_of p) \ s).nonempty, - { ext i, - simp }, - refine (hp.diff $ (set.finite_lt_nat _).bUnion _).nonempty, - exact λ k h, set.finite_le_nat _, + rw [nth_eq_order_iso_of_nat hf], + haveI := hf.to_subtype, + exact nat.subtype.coe_comp_of_nat_range end -lemma nth_mem_of_infinite (hp : (set_of p).infinite) (n : ℕ) : p (nth p n) := -(nth_mem_of_infinite_aux p hp n).1 +theorem nth_mem_of_infinite (hf : (set_of p).infinite) (n : ℕ) : p (nth p n) := +set.range_subset_iff.1 (range_nth_of_infinite hf).le n -lemma nth_strict_mono (hp : (set_of p).infinite) : strict_mono (nth p) := -λ a b, (nth_mem_of_infinite_aux p hp b).2 _ +/-! +### Lemmas that work for finite and infinite sets +-/ -lemma nth_injective_of_infinite (hp : (set_of p).infinite) : function.injective (nth p) := +theorem exists_lt_card_nth_eq {x} (h : p x) : + ∃ n, (∀ hf : (set_of p).finite, n < hf.to_finset.card) ∧ nth p n = x := begin - intros m n h, - wlog h' : m ≤ n, - { exact (this p hp h.symm (le_of_not_le h')).symm }, - rw le_iff_lt_or_eq at h', - obtain (h' | rfl) := h', - { simpa [h] using nth_strict_mono p hp h' }, - { refl }, + refine (set_of p).finite_or_infinite.elim (λ hf, _) (λ hf, _), + { rcases exists_lt_card_finite_nth_eq hf h with ⟨n, hn, hx⟩, + exact ⟨n, λ hf', hn, hx⟩ }, + { rw [← @set.mem_set_of_eq _ _ p, ← range_nth_of_infinite hf] at h, + rcases h with ⟨n, hx⟩, + exact ⟨n, λ hf', absurd hf' hf, hx⟩ } end -lemma nth_monotone (hp : (set_of p).infinite) : monotone (nth p) := -(nth_strict_mono p hp).monotone +theorem subset_range_nth : set_of p ⊆ set.range (nth p) := +λ x (hx : p x), let ⟨n, _, hn⟩ := exists_lt_card_nth_eq hx in ⟨n, hn⟩ + +theorem range_nth_subset : set.range (nth p) ⊆ insert 0 (set_of p) := +(set_of p).finite_or_infinite.elim (λ h, (range_nth_of_finite h).subset) + (λ h, (range_nth_of_infinite h).trans_subset (set.subset_insert _ _)) + +theorem nth_mem (n : ℕ) (h : ∀ hf : (set_of p).finite, n < hf.to_finset.card) : p (nth p n) := +(set_of p).finite_or_infinite.elim (λ hf, nth_mem_of_lt_card hf (h hf)) + (λ h, nth_mem_of_infinite h n) + +theorem nth_lt_nth' {m n : ℕ} (hlt : m < n) (h : ∀ hf : (set_of p).finite, n < hf.to_finset.card) : + nth p m < nth p n := +(set_of p).finite_or_infinite.elim (λ hf, nth_lt_nth_of_lt_card hf hlt (h _)) + (λ hf, (nth_lt_nth hf).2 hlt) + +theorem nth_le_nth' {m n : ℕ} (hle : m ≤ n) (h : ∀ hf : (set_of p).finite, n < hf.to_finset.card) : + nth p m ≤ nth p n := +(set_of p).finite_or_infinite.elim (λ hf, nth_le_nth_of_lt_card hf hle (h _)) + (λ hf, (nth_le_nth hf).2 hle) + +theorem le_nth {n : ℕ} (h : ∀ hf : (set_of p).finite, n < hf.to_finset.card) : n ≤ nth p n := +(set_of p).finite_or_infinite.elim + (λ hf, ((nth_strict_mono_on hf).mono $ set.Iic_subset_Iio.2 (h _)).Iic_id_le _ le_rfl) + (λ hf, (nth_strict_mono hf).id_le _) + +theorem is_least_nth {n} (h : ∀ hf : (set_of p).finite, n < hf.to_finset.card) : + is_least { i | p i ∧ ∀ k < n, nth p k < i } (nth p n) := +⟨⟨nth_mem n h, λ k hk, nth_lt_nth' hk h⟩, λ x hx, let ⟨k, hk, hkx⟩ := exists_lt_card_nth_eq hx.1 + in (lt_or_le k n).elim (λ hlt, absurd hkx (hx.2 _ hlt).ne) (λ hle, hkx ▸ nth_le_nth' hle hk)⟩ + +theorem is_least_nth_of_lt_card {n : ℕ} (hf : (set_of p).finite) (hn : n < hf.to_finset.card) : + is_least { i | p i ∧ ∀ k < n, nth p k < i } (nth p n) := +is_least_nth $ λ _, hn + +theorem is_least_nth_of_infinite (hf : (set_of p).infinite) (n : ℕ) : + is_least { i | p i ∧ ∀ k < n, nth p k < i } (nth p n) := +is_least_nth $ λ h, absurd h hf + +/-- An alternative recursive definition of `nat.nth`: `nat.nth s n` is the infimum of `x ∈ s` such +that `nat.nth s k < x` for all `k < n`, if this set is nonempty. We do not assume that the set is +nonempty because we use the same "garbage value" `0` both for `Inf` on `ℕ` and for `nat.nth s n` for +`n ≥ card s`. -/ +lemma nth_eq_Inf (p : ℕ → Prop) (n : ℕ) : nth p n = Inf {x | p x ∧ ∀ k < n, nth p k < x} := +begin + by_cases hn : ∀ hf : (set_of p).finite, n < hf.to_finset.card, + { exact (is_least_nth hn).cInf_eq.symm }, + { push_neg at hn, + rcases hn with ⟨hf, hn⟩, + rw [nth_of_card_le _ hn], + refine ((congr_arg Inf $ set.eq_empty_of_forall_not_mem $ λ k hk, _).trans Inf_empty).symm, + rcases exists_lt_card_nth_eq hk.1 with ⟨k, hlt, rfl⟩, + exact (hk.2 _ ((hlt hf).trans_le hn)).false } +end + +lemma nth_zero : nth p 0 = Inf (set_of p) := by { rw nth_eq_Inf, simp } + +@[simp] lemma nth_zero_of_zero (h : p 0) : nth p 0 = 0 := by simp [nth_zero, h] + +lemma nth_zero_of_exists [decidable_pred p] (h : ∃ n, p n) : nth p 0 = nat.find h := +by { rw [nth_zero], convert nat.Inf_def h } -lemma nth_mono_of_finite {a b : ℕ} (hp : (set_of p).finite) (hb : b < hp.to_finset.card) - (hab : a ≤ b) : nth p a ≤ nth p b := +lemma nth_eq_zero {n} : + nth p n = 0 ↔ p 0 ∧ n = 0 ∨ ∃ hf : (set_of p).finite, hf.to_finset.card ≤ n := begin - obtain rfl | h := hab.eq_or_lt, - { exact le_rfl }, - { exact (nth_strict_mono_of_finite p hp hb h).le } + refine ⟨λ h, _, _⟩, + { simp only [or_iff_not_imp_right, not_exists, not_le], + exact λ hn, ⟨h ▸ nth_mem _ hn, nonpos_iff_eq_zero.1 $ h ▸ le_nth hn⟩ }, + { rintro (⟨h₀, rfl⟩ | ⟨hf, hle⟩), + exacts [nth_zero_of_zero h₀, nth_of_card_le hf hle] } end -lemma le_nth_of_lt_nth_succ_finite {k a : ℕ} (hp : (set_of p).finite) - (hlt : k.succ < hp.to_finset.card) (h : a < nth p k.succ) (ha : p a) : - a ≤ nth p k := +lemma nth_eq_zero_mono (h₀ : ¬p 0) {a b : ℕ} (hab : a ≤ b) (ha : nth p a = 0) : + nth p b = 0 := begin - by_contra' hak, - refine h.not_le _, - rw nth, - apply nat.Inf_le, - refine ⟨ha, λ n hn, lt_of_le_of_lt _ hak⟩, - exact nth_mono_of_finite p hp (k.le_succ.trans_lt hlt) (le_of_lt_succ hn), + simp only [nth_eq_zero, h₀, false_and, false_or] at ha ⊢, + exact ha.imp (λ hf hle, hle.trans hab) end -lemma le_nth_of_lt_nth_succ_infinite {k a : ℕ} (hp : (set_of p).infinite) - (h : a < nth p k.succ) (ha : p a) : +lemma le_nth_of_lt_nth_succ {k a : ℕ} (h : a < nth p (k + 1)) (ha : p a) : a ≤ nth p k := begin - by_contra' hak, - refine h.not_le _, - rw nth, - apply nat.Inf_le, - exact ⟨ha, λ n hn, (nth_monotone p hp (le_of_lt_succ hn)).trans_lt hak⟩, + cases (set_of p).finite_or_infinite with hf hf, + { rcases exists_lt_card_finite_nth_eq hf ha with ⟨n, hn, rfl⟩, + cases lt_or_le (k + 1) hf.to_finset.card with hk hk, + { rwa [(nth_strict_mono_on hf).lt_iff_lt hn hk, lt_succ_iff, + ← (nth_strict_mono_on hf).le_iff_le hn (k.lt_succ_self.trans hk)] at h }, + { rw [nth_of_card_le _ hk] at h, + exact absurd h (zero_le _).not_lt } }, + { rcases subset_range_nth ha with ⟨n, rfl⟩, + rwa [nth_lt_nth hf, lt_succ_iff, ← nth_le_nth hf] at h } end section count -variables [decidable_pred p] +variables (p) [decidable_pred p] @[simp] lemma count_nth_zero : count p (nth p 0) = 0 := begin - rw [count_eq_card_filter_range, finset.card_eq_zero, nth_zero], - ext a, - simp_rw [not_mem_empty, mem_filter, mem_range, iff_false], - rintro ⟨ha, hp⟩, - exact ha.not_le (nat.Inf_le hp), + rw [count_eq_card_filter_range, card_eq_zero, filter_eq_empty_iff, nth_zero], + exact λ n h₁ h₂, (mem_range.1 h₁).not_le (nat.Inf_le h₂) end -lemma filter_range_nth_eq_insert_of_finite (hp : (set_of p).finite) {k : ℕ} - (hlt : k.succ < hp.to_finset.card) : - finset.filter p (finset.range (nth p k.succ)) = - insert (nth p k) (finset.filter p (finset.range (nth p k))) := +lemma filter_range_nth_subset_insert (k : ℕ) : + (range (nth p (k + 1))).filter p ⊆ insert (nth p k) ((range (nth p k)).filter p) := begin - ext a, - simp_rw [mem_insert, mem_filter, mem_range], - split, - { rintro ⟨ha, hpa⟩, - refine or_iff_not_imp_left.mpr (λ h, ⟨lt_of_le_of_ne _ h, hpa⟩), - exact le_nth_of_lt_nth_succ_finite p hp hlt ha hpa }, - { rintro (ha | ⟨ha, hpa⟩), - { rw ha, - refine ⟨nth_strict_mono_of_finite p hp hlt (lt_add_one _), _⟩, - apply nth_mem_of_lt_card_finite p hp, - exact (k.le_succ).trans_lt hlt }, - refine ⟨ha.trans _, hpa⟩, - exact nth_strict_mono_of_finite p hp hlt (lt_add_one _) } + intros a ha, + simp only [mem_insert, mem_filter, mem_range] at ha ⊢, + exact (le_nth_of_lt_nth_succ ha.1 ha.2).eq_or_lt.imp_right (λ h, ⟨h, ha.2⟩) end -lemma count_nth_of_lt_card_finite {n : ℕ} (hp : (set_of p).finite) - (hlt : n < hp.to_finset.card) : count p (nth p n) = n := +variable {p} + +lemma filter_range_nth_eq_insert {k : ℕ} + (hlt : ∀ hf : (set_of p).finite, k + 1 < hf.to_finset.card) : + (range (nth p (k + 1))).filter p = insert (nth p k) ((range (nth p k)).filter p) := begin - induction n with k hk, - { exact count_nth_zero _ }, - { rw [count_eq_card_filter_range, filter_range_nth_eq_insert_of_finite p hp hlt, - finset.card_insert_of_not_mem, ←count_eq_card_filter_range, hk (lt_of_succ_lt hlt)], - simp, }, + refine (filter_range_nth_subset_insert p k).antisymm (λ a ha, _), + simp only [mem_insert, mem_filter, mem_range] at ha ⊢, + have : nth p k < nth p (k + 1) := nth_lt_nth' k.lt_succ_self hlt, + rcases ha with (rfl | ⟨hlt, hpa⟩), + { exact ⟨this, nth_mem _ (λ hf, k.lt_succ_self.trans (hlt hf))⟩ }, + { exact ⟨hlt.trans this, hpa⟩ }, end +lemma filter_range_nth_eq_insert_of_finite (hf : (set_of p).finite) {k : ℕ} + (hlt : k + 1 < hf.to_finset.card) : + (range (nth p (k + 1))).filter p = insert (nth p k) ((range (nth p k)).filter p) := +filter_range_nth_eq_insert $ λ _, hlt + lemma filter_range_nth_eq_insert_of_infinite (hp : (set_of p).infinite) (k : ℕ) : - (finset.range (nth p k.succ)).filter p = insert (nth p k) ((finset.range (nth p k)).filter p) := -begin - ext a, - simp_rw [mem_insert, mem_filter, mem_range], - split, - { rintro ⟨ha, hpa⟩, - rw nth at ha, - refine or_iff_not_imp_left.mpr (λ hne, ⟨(le_of_not_lt $ λ h, _).lt_of_ne hne, hpa⟩), - exact ha.not_le (nat.Inf_le ⟨hpa, λ b hb, (nth_monotone p hp (le_of_lt_succ hb)).trans_lt h⟩) }, - { rintro (rfl | ⟨ha, hpa⟩), - { exact ⟨nth_strict_mono p hp (lt_succ_self k), nth_mem_of_infinite p hp _⟩ }, - { exact ⟨ha.trans (nth_strict_mono p hp (lt_succ_self k)), hpa⟩ } } -end + (range (nth p (k + 1))).filter p = insert (nth p k) ((range (nth p k)).filter p) := +filter_range_nth_eq_insert $ λ hf, absurd hf hp -lemma count_nth_of_infinite (hp : (set_of p).infinite) (n : ℕ) : count p (nth p n) = n := +lemma count_nth {n : ℕ} (hn : ∀ hf : (set_of p).finite, n < hf.to_finset.card) : + count p (nth p n) = n := begin - induction n with k hk, + induction n with k ihk, { exact count_nth_zero _ }, - { rw [count_eq_card_filter_range, filter_range_nth_eq_insert_of_infinite p hp, - finset.card_insert_of_not_mem, ←count_eq_card_filter_range, hk], - simp, }, + { rw [count_eq_card_filter_range, filter_range_nth_eq_insert hn, card_insert_of_not_mem, + ←count_eq_card_filter_range, ihk (λ hf, lt_of_succ_lt (hn hf))], + simp } end +lemma count_nth_of_lt_card_finite {n : ℕ} (hp : (set_of p).finite) + (hlt : n < hp.to_finset.card) : count p (nth p n) = n := +count_nth $ λ _, hlt + +lemma count_nth_of_infinite (hp : (set_of p).infinite) (n : ℕ) : count p (nth p n) = n := +count_nth $ λ hf, absurd hf hp + +lemma count_nth_succ {n : ℕ} (hn : ∀ hf : (set_of p).finite, n < hf.to_finset.card) : + count p (nth p n + 1) = n + 1 := +by rw [count_succ, count_nth hn, if_pos (nth_mem _ hn)] + @[simp] lemma nth_count {n : ℕ} (hpn : p n) : nth p (count p n) = n := -begin - obtain hp | hp := em (set_of p).finite, - { refine count_injective _ hpn _, - { apply nth_mem_of_lt_card_finite p hp, - exact count_lt_card hp hpn }, - { exact count_nth_of_lt_card_finite _ _ (count_lt_card hp hpn) } }, - { apply count_injective (nth_mem_of_infinite _ hp _) hpn, - apply count_nth_of_infinite p hp } -end +have ∀ hf : (set_of p).finite, count p n < hf.to_finset.card, + from λ hf, count_lt_card hf hpn, +count_injective (nth_mem _ this) hpn (count_nth this) -lemma nth_count_eq_Inf {n : ℕ} : nth p (count p n) = Inf {i : ℕ | p i ∧ n ≤ i} := +lemma nth_lt_of_lt_count {n k : ℕ} (h : k < count p n) : nth p k < n := begin - rw nth, - congr, - ext a, - simp only [set.mem_set_of_eq, and.congr_right_iff], - intro hpa, - refine ⟨λ h, _, λ hn k hk, lt_of_lt_of_le _ hn⟩, - { by_contra ha, - simp only [not_le] at ha, - have hn : nth p (count p a) < a := h _ (count_strict_mono hpa ha), - rwa [nth_count p hpa, lt_self_iff_false] at hn }, - { apply (count_monotone p).reflect_lt, - convert hk, - obtain hp | hp : (set_of p).finite ∨ (set_of p).infinite := em (set_of p).finite, - { rw count_nth_of_lt_card_finite _ hp, - exact hk.trans ((count_monotone _ hn).trans_lt (count_lt_card hp hpa)) }, - { apply count_nth_of_infinite p hp } } + refine (count_monotone p).reflect_lt _, + rwa [count_nth], + exact λ hf, h.trans_le (count_le_card hf n) end -lemma nth_count_le (hp : (set_of p).infinite) (n : ℕ) : n ≤ nth p (count p n) := -begin - rw nth_count_eq_Inf, - suffices h : Inf {i : ℕ | p i ∧ n ≤ i} ∈ {i : ℕ | p i ∧ n ≤ i}, - { exact h.2 }, - apply Inf_mem, - obtain ⟨m, hp, hn⟩ := hp.exists_gt n, - exact ⟨m, hp, hn.le⟩ -end +lemma le_nth_of_count_le {n k : ℕ} (h : n ≤ nth p k) : count p n ≤ k := +not_lt.1 $ λ hlt, h.not_lt $ nth_lt_of_lt_count hlt -lemma count_nth_gc (hp : (set_of p).infinite) : galois_connection (count p) (nth p) := +variable (p) + +lemma nth_count_eq_Inf (n : ℕ) : nth p (count p n) = Inf {i : ℕ | p i ∧ n ≤ i} := begin - rintro x y, - rw [nth, le_cInf_iff ⟨0, λ _ _, nat.zero_le _⟩ ⟨nth p y, nth_mem_of_infinite_aux p hp y⟩], - dsimp, - refine ⟨_, λ h, _⟩, - { rintro hy n ⟨hn, h⟩, - obtain hy' | rfl := hy.lt_or_eq, - { exact (nth_count_le p hp x).trans (h (count p x) hy').le }, - { specialize h (count p n), - replace hn : nth p (count p n) = n := nth_count _ hn, - replace h : count p x ≤ count p n := by rwa [hn, lt_self_iff_false, imp_false, not_lt] at h, - refine (nth_count_le p hp x).trans _, - rw ← hn, - exact nth_monotone p hp h }, }, - { rw ←count_nth_of_infinite p hp y, - exact count_monotone _ (h (nth p y) ⟨nth_mem_of_infinite p hp y, - λ k hk, nth_strict_mono p hp hk⟩) } + refine (nth_eq_Inf _ _).trans (congr_arg Inf _), + refine set.ext (λ a, and_congr_right $ λ hpa, _), + refine ⟨λ h, not_lt.1 (λ ha, _), λ hn k hk, lt_of_lt_of_le (nth_lt_of_lt_count hk) hn⟩, + have hn : nth p (count p a) < a := h _ (count_strict_mono hpa ha), + rwa [nth_count hpa, lt_self_iff_false] at hn end -lemma count_le_iff_le_nth (hp : (set_of p).infinite) {a b : ℕ} : - count p a ≤ b ↔ a ≤ nth p b := count_nth_gc p hp _ _ +variable {p} -lemma lt_nth_iff_count_lt (hp : (set_of p).infinite) {a b : ℕ} : - a < count p b ↔ nth p a < b := lt_iff_lt_of_le_iff_le $ count_le_iff_le_nth p hp +lemma le_nth_count' {n : ℕ} (hpn : ∃ k, p k ∧ n ≤ k) : n ≤ nth p (count p n) := +(le_cInf hpn $ λ k, and.right).trans (nth_count_eq_Inf p n).ge -lemma nth_lt_of_lt_count (n k : ℕ) (h : k < count p n) : nth p k < n := -begin - obtain hp | hp := em (set_of p).finite, - { refine (count_monotone p).reflect_lt _, - rwa count_nth_of_lt_card_finite p hp, - refine h.trans_le _, - rw count_eq_card_filter_range, - exact finset.card_le_of_subset (λ x hx, hp.mem_to_finset.2 (mem_filter.1 hx).2) }, - { rwa ← lt_nth_iff_count_lt _ hp } -end +lemma le_nth_count (hp : (set_of p).infinite) (n : ℕ) : n ≤ nth p (count p n) := +let ⟨m, hp, hn⟩ := hp.exists_gt n in le_nth_count' ⟨m, hp, hn.le⟩ -lemma le_nth_of_count_le (n k : ℕ) (h: n ≤ nth p k) : count p n ≤ k := -begin - by_contra hc, - apply not_lt.mpr h, - apply nth_lt_of_lt_count, - simpa using hc -end +/-- If a predicate `p : ℕ → Prop` is true for infinitely many numbers, then `nat.count p` and +`nat.nth p` form a Galois insertion. -/ +noncomputable def gi_count_nth (hp : (set_of p).infinite) : galois_insertion (count p) (nth p) := +galois_insertion.monotone_intro (nth_monotone hp) (count_monotone p) + (le_nth_count hp) (count_nth_of_infinite hp) -end count +lemma gc_count_nth (hp : (set_of p).infinite) : galois_connection (count p) (nth p) := +(gi_count_nth hp).gc -lemma nth_zero_of_nth_zero (h₀ : ¬p 0) {a b : ℕ} (hab : a ≤ b) (ha : nth p a = 0) : - nth p b = 0 := -begin - rw [nth, Inf_eq_zero] at ⊢ ha, - cases ha, - { exact (h₀ ha.1).elim }, - { refine or.inr (set.eq_empty_of_subset_empty $ λ x hx, _), - rw ←ha, - exact ⟨hx.1, λ k hk, hx.2 k $ hk.trans_le hab⟩ } -end +lemma count_le_iff_le_nth (hp : (set_of p).infinite) {a b : ℕ} : + count p a ≤ b ↔ a ≤ nth p b := gc_count_nth hp _ _ -/-- When `p` is true infinitely often, `nth` agrees with `nat.subtype.order_iso_of_nat`. -/ -lemma nth_eq_order_iso_of_nat (i : infinite (set_of p)) (n : ℕ) : - nth p n = nat.subtype.order_iso_of_nat (set_of p) n := -begin - classical, - have hi := set.infinite_coe_iff.mp i, - induction n with k hk; - simp only [subtype.order_iso_of_nat_apply, subtype.of_nat, nat_zero_eq_zero], - { rw [nat.subtype.coe_bot, nth_zero_of_exists], }, - { simp only [nat.subtype.succ, set.mem_set_of_eq, subtype.coe_mk, subtype.val_eq_coe], - rw [subtype.order_iso_of_nat_apply] at hk, - set b := nth p k.succ - nth p k - 1 with hb, - replace hb : p (↑(subtype.of_nat (set_of p) k) + b + 1), - { rw [hb, ←hk, tsub_right_comm], - have hn11: nth p k.succ - 1 + 1 = nth p k.succ, - { rw tsub_add_cancel_iff_le, - exact succ_le_of_lt (pos_of_gt (nth_strict_mono p hi (lt_add_one k))), }, - rw add_tsub_cancel_of_le, - { rw hn11, - apply nth_mem_of_infinite p hi }, - { rw [← lt_succ_iff, ← nat.add_one, hn11], - apply nth_strict_mono p hi, - exact lt_add_one k } }, - have H : (∃ n: ℕ , p (↑(subtype.of_nat (set_of p) k) + n + 1)) := ⟨b, hb⟩, - set t := nat.find H with ht, - obtain ⟨hp, hmin⟩ := (nat.find_eq_iff _).mp ht, - rw [←ht, ←hk] at hp hmin ⊢, - rw [nth, Inf_def ⟨_, nth_mem_of_infinite_aux p hi k.succ⟩, nat.find_eq_iff], - refine ⟨⟨by convert hp, λ r hr, _⟩, λ n hn, _⟩, - { rw lt_succ_iff at ⊢ hr, - exact (nth_monotone p hi hr).trans (by simp) }, - simp only [exists_prop, not_and, not_lt, set.mem_set_of_eq, not_forall], - refine λ hpn, ⟨k, lt_add_one k, _⟩, - by_contra' hlt, - replace hn : n - nth p k - 1 < t, - { rw tsub_lt_iff_left, - { rw tsub_lt_iff_left hlt.le, - convert hn using 1, - ac_refl }, - exact le_tsub_of_add_le_left (succ_le_of_lt hlt) }, - refine hmin (n - nth p k - 1) hn _, - convert hpn, - have hn11 : n - 1 + 1 = n := nat.sub_add_cancel (pos_of_gt hlt), - rwa [tsub_right_comm, add_tsub_cancel_of_le], - rwa [←hn11, lt_succ_iff] at hlt }, -end +lemma lt_nth_iff_count_lt (hp : (set_of p).infinite) {a b : ℕ} : + a < count p b ↔ nth p a < b := (gc_count_nth hp).lt_iff_lt + +end count end nat diff --git a/src/data/set/finite.lean b/src/data/set/finite.lean index d9918a209a506..67eb58978f762 100644 --- a/src/data/set/finite.lean +++ b/src/data/set/finite.lean @@ -98,6 +98,14 @@ h.nonempty_fintype.some protected noncomputable def finite.to_finset {s : set α} (h : s.finite) : finset α := @set.to_finset _ _ h.fintype +theorem finite.to_finset_eq_to_finset {s : set α} [fintype s] (h : s.finite) : + h.to_finset = s.to_finset := +by { rw [finite.to_finset], congr } + +@[simp] +theorem to_finite_to_finset (s : set α) [fintype s] : s.to_finite.to_finset = s.to_finset := +s.to_finite.to_finset_eq_to_finset + theorem finite.exists_finset {s : set α} (h : s.finite) : ∃ s' : finset α, ∀ a : α, a ∈ s' ↔ a ∈ s := by { casesI h, exact ⟨s.to_finset, λ _, mem_to_finset⟩ } diff --git a/src/number_theory/prime_counting.lean b/src/number_theory/prime_counting.lean index b531b97019305..d811b98cd28f0 100644 --- a/src/number_theory/prime_counting.lean +++ b/src/number_theory/prime_counting.lean @@ -56,10 +56,10 @@ lemma monotone_prime_counting : monotone prime_counting := monotone_prime_counting'.comp (monotone_id.add_const _) @[simp] lemma prime_counting'_nth_eq (n : ℕ) : π' (nth prime n) = n := -count_nth_of_infinite _ infinite_set_of_prime _ +count_nth_of_infinite infinite_set_of_prime _ @[simp] lemma prime_nth_prime (n : ℕ) : prime (nth prime n) := -nth_mem_of_infinite _ infinite_set_of_prime _ +nth_mem_of_infinite infinite_set_of_prime _ /-- A linear upper bound on the size of the `prime_counting'` function -/ lemma prime_counting'_add_le {a k : ℕ} (h0 : 0 < a) (h1 : a < k) (n : ℕ) : From 087c325ae0ab42dbdd5dee55bc37d3d5a0bf2197 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Mon, 12 Jun 2023 06:38:55 +0000 Subject: [PATCH 1173/1291] chore(*): add mathlib4 synchronization comments (#19176) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.category.Module.filtered_colimits` * `analysis.normed_space.algebra` * `analysis.normed_space.spectrum` * `analysis.normed_space.star.gelfand_duality` * `analysis.normed_space.star.spectrum` * `category_theory.sites.surjective` * `measure_theory.measure.hausdorff` * `number_theory.dioph` * `number_theory.wilson` * `ring_theory.perfection` --- src/algebra/category/Module/filtered_colimits.lean | 3 +++ src/analysis/normed_space/algebra.lean | 3 +++ src/analysis/normed_space/spectrum.lean | 3 +++ src/analysis/normed_space/star/gelfand_duality.lean | 3 +++ src/analysis/normed_space/star/spectrum.lean | 3 +++ src/category_theory/sites/surjective.lean | 3 +++ src/measure_theory/measure/hausdorff.lean | 3 +++ src/number_theory/dioph.lean | 3 +++ src/number_theory/wilson.lean | 3 +++ src/ring_theory/perfection.lean | 3 +++ 10 files changed, 30 insertions(+) diff --git a/src/algebra/category/Module/filtered_colimits.lean b/src/algebra/category/Module/filtered_colimits.lean index 4835fbaefa044..01cb3e6a786ef 100644 --- a/src/algebra/category/Module/filtered_colimits.lean +++ b/src/algebra/category/Module/filtered_colimits.lean @@ -9,6 +9,9 @@ import algebra.category.Module.basic /-! # The forgetful functor from `R`-modules preserves filtered colimits. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Forgetful functors from algebraic categories usually don't preserve colimits. However, they tend to preserve _filtered_ colimits. diff --git a/src/analysis/normed_space/algebra.lean b/src/analysis/normed_space/algebra.lean index 62336f754f7b8..b3ae7a3ec2ba8 100644 --- a/src/analysis/normed_space/algebra.lean +++ b/src/analysis/normed_space/algebra.lean @@ -11,6 +11,9 @@ import analysis.normed_space.spectrum /-! # Normed algebras +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains basic facts about normed algebras. ## Main results diff --git a/src/analysis/normed_space/spectrum.lean b/src/analysis/normed_space/spectrum.lean index 8fcc2fedf5890..09dbf5e2d5006 100644 --- a/src/analysis/normed_space/spectrum.lean +++ b/src/analysis/normed_space/spectrum.lean @@ -12,6 +12,9 @@ import analysis.normed_space.exponential /-! # The spectrum of elements in a complete normed algebra +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains the basic theory for the resolvent and spectrum of a Banach algebra. ## Main definitions diff --git a/src/analysis/normed_space/star/gelfand_duality.lean b/src/analysis/normed_space/star/gelfand_duality.lean index 236d6b2b7c563..58777a5d3fd29 100644 --- a/src/analysis/normed_space/star/gelfand_duality.lean +++ b/src/analysis/normed_space/star/gelfand_duality.lean @@ -14,6 +14,9 @@ import topology.continuous_function.stone_weierstrass /-! # Gelfand Duality +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The `gelfand_transform` is an algebra homomorphism from a topological `𝕜`-algebra `A` to `C(character_space 𝕜 A, 𝕜)`. In the case where `A` is a commutative complex Banach algebra, then the Gelfand transform is actually spectrum-preserving (`spectrum.gelfand_transform_eq`). Moreover, diff --git a/src/analysis/normed_space/star/spectrum.lean b/src/analysis/normed_space/star/spectrum.lean index c9cf202ad89d5..2680b98fdefa0 100644 --- a/src/analysis/normed_space/star/spectrum.lean +++ b/src/analysis/normed_space/star/spectrum.lean @@ -9,6 +9,9 @@ import analysis.special_functions.exponential import algebra.star.star_alg_hom /-! # Spectral properties in C⋆-algebras + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. In this file, we establish various properties related to the spectrum of elements in C⋆-algebras. -/ diff --git a/src/category_theory/sites/surjective.lean b/src/category_theory/sites/surjective.lean index 93c68c5866ca0..ecaac4b81c737 100644 --- a/src/category_theory/sites/surjective.lean +++ b/src/category_theory/sites/surjective.lean @@ -10,6 +10,9 @@ import category_theory.sites.compatible_sheafification # Locally surjective morphisms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions - `is_locally_surjective` : A morphism of presheaves valued in a concrete category is locally diff --git a/src/measure_theory/measure/hausdorff.lean b/src/measure_theory/measure/hausdorff.lean index 75c022de2443c..40b9e7863e06f 100644 --- a/src/measure_theory/measure/hausdorff.lean +++ b/src/measure_theory/measure/hausdorff.lean @@ -13,6 +13,9 @@ import topology.metric_space.metric_separated /-! # Hausdorff measure and metric (outer) measures +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define the `d`-dimensional Hausdorff measure on an (extended) metric space `X` and the Hausdorff dimension of a set in an (extended) metric space. Let `μ d δ` be the maximal outer measure such that `μ d δ s ≤ (emetric.diam s) ^ d` for every set of diameter less than `δ`. Then diff --git a/src/number_theory/dioph.lean b/src/number_theory/dioph.lean index 5b5d2b23291d1..7cb13e17d1d05 100644 --- a/src/number_theory/dioph.lean +++ b/src/number_theory/dioph.lean @@ -11,6 +11,9 @@ import number_theory.pell_matiyasevic /-! # Diophantine functions and Matiyasevic's theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Hilbert's tenth problem asked whether there exists an algorithm which for a given integer polynomial determines whether this polynomial has integer solutions. It was answered in the negative in 1970, the final step being completed by Matiyasevic who showed that the power function is Diophantine. diff --git a/src/number_theory/wilson.lean b/src/number_theory/wilson.lean index b253a9dbb3c90..c077fc3d2fbec 100644 --- a/src/number_theory/wilson.lean +++ b/src/number_theory/wilson.lean @@ -8,6 +8,9 @@ import field_theory.finite.basic /-! # Wilson's theorem. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains a proof of Wilson's theorem. The heavy lifting is mostly done by the previous `wilsons_lemma`, diff --git a/src/ring_theory/perfection.lean b/src/ring_theory/perfection.lean index 8357e7a373fd7..519d14d808be5 100644 --- a/src/ring_theory/perfection.lean +++ b/src/ring_theory/perfection.lean @@ -17,6 +17,9 @@ import ring_theory.valuation.integers /-! # Ring Perfection and Tilt +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define the perfection of a ring of characteristic p, and the tilt of a field given a valuation to `ℝ≥0`. From d8bbb04e2d2a44596798a9207ceefc0fb236e41e Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Mon, 12 Jun 2023 09:33:15 +0000 Subject: [PATCH 1174/1291] chore(measure_theory/function/conditional_expectation): split `measure_theory.function.conditional_expectation.basic` (#19177) --- .../ae_measurable.lean | 698 +++++++ .../conditional_expectation/basic.lean | 1859 +---------------- .../conditional_expectation/condexp_L1.lean | 567 +++++ .../conditional_expectation/condexp_L2.lean | 518 +++++ .../conditional_expectation/unique.lean | 213 ++ 5 files changed, 2005 insertions(+), 1850 deletions(-) create mode 100644 src/measure_theory/function/conditional_expectation/ae_measurable.lean create mode 100644 src/measure_theory/function/conditional_expectation/condexp_L1.lean create mode 100644 src/measure_theory/function/conditional_expectation/condexp_L2.lean create mode 100644 src/measure_theory/function/conditional_expectation/unique.lean diff --git a/src/measure_theory/function/conditional_expectation/ae_measurable.lean b/src/measure_theory/function/conditional_expectation/ae_measurable.lean new file mode 100644 index 0000000000000..ed446146e1a82 --- /dev/null +++ b/src/measure_theory/function/conditional_expectation/ae_measurable.lean @@ -0,0 +1,698 @@ +/- +Copyright (c) 2021 Rémy Degenne. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Rémy Degenne +-/ +import measure_theory.function.l2_space +import measure_theory.function.strongly_measurable.lp + +/-! # Functions a.e. measurable with respect to a sub-σ-algebra + +A function `f` verifies `ae_strongly_measurable' m f μ` if it is `μ`-a.e. equal to +an `m`-strongly measurable function. This is similar to `ae_strongly_measurable`, but the +`measurable_space` structures used for the measurability statement and for the measure are +different. + +We define `Lp_meas F 𝕜 m p μ`, the subspace of `Lp F p μ` containing functions `f` verifying +`ae_strongly_measurable' m f μ`, i.e. functions which are `μ`-a.e. equal to an `m`-strongly +measurable function. + +## Main statements + +We define an `isometry_equiv` between `Lp_meas_subgroup` and the `Lp` space corresponding to the +measure `μ.trim hm`. As a consequence, the completeness of `Lp` implies completeness of `Lp_meas`. + +`Lp.induction_strongly_measurable` (see also `mem_ℒp.induction_strongly_measurable`): +To prove something for an `Lp` function a.e. strongly measurable with respect to a +sub-σ-algebra `m` in a normed space, it suffices to show that +* the property holds for (multiples of) characteristic functions which are measurable w.r.t. `m`; +* is closed under addition; +* the set of functions in `Lp` strongly measurable w.r.t. `m` for which the property holds is + closed. + +-/ + +open topological_space filter +open_locale ennreal measure_theory + +namespace measure_theory + +/-- A function `f` verifies `ae_strongly_measurable' m f μ` if it is `μ`-a.e. equal to +an `m`-strongly measurable function. This is similar to `ae_strongly_measurable`, but the +`measurable_space` structures used for the measurability statement and for the measure are +different. -/ +def ae_strongly_measurable' {α β} [topological_space β] + (m : measurable_space α) {m0 : measurable_space α} + (f : α → β) (μ : measure α) : Prop := +∃ g : α → β, strongly_measurable[m] g ∧ f =ᵐ[μ] g + +namespace ae_strongly_measurable' + +variables {α β 𝕜 : Type*} {m m0 : measurable_space α} {μ : measure α} + [topological_space β] {f g : α → β} + +lemma congr (hf : ae_strongly_measurable' m f μ) (hfg : f =ᵐ[μ] g) : + ae_strongly_measurable' m g μ := +by { obtain ⟨f', hf'_meas, hff'⟩ := hf, exact ⟨f', hf'_meas, hfg.symm.trans hff'⟩, } + +lemma add [has_add β] [has_continuous_add β] (hf : ae_strongly_measurable' m f μ) + (hg : ae_strongly_measurable' m g μ) : + ae_strongly_measurable' m (f+g) μ := +begin + rcases hf with ⟨f', h_f'_meas, hff'⟩, + rcases hg with ⟨g', h_g'_meas, hgg'⟩, + exact ⟨f' + g', h_f'_meas.add h_g'_meas, hff'.add hgg'⟩, +end + +lemma neg [add_group β] [topological_add_group β] + {f : α → β} (hfm : ae_strongly_measurable' m f μ) : + ae_strongly_measurable' m (-f) μ := +begin + rcases hfm with ⟨f', hf'_meas, hf_ae⟩, + refine ⟨-f', hf'_meas.neg, hf_ae.mono (λ x hx, _)⟩, + simp_rw pi.neg_apply, + rw hx, +end + +lemma sub [add_group β] [topological_add_group β] {f g : α → β} + (hfm : ae_strongly_measurable' m f μ) (hgm : ae_strongly_measurable' m g μ) : + ae_strongly_measurable' m (f - g) μ := +begin + rcases hfm with ⟨f', hf'_meas, hf_ae⟩, + rcases hgm with ⟨g', hg'_meas, hg_ae⟩, + refine ⟨f'-g', hf'_meas.sub hg'_meas, hf_ae.mp (hg_ae.mono (λ x hx1 hx2, _))⟩, + simp_rw pi.sub_apply, + rw [hx1, hx2], +end + +lemma const_smul [has_smul 𝕜 β] [has_continuous_const_smul 𝕜 β] + (c : 𝕜) (hf : ae_strongly_measurable' m f μ) : + ae_strongly_measurable' m (c • f) μ := +begin + rcases hf with ⟨f', h_f'_meas, hff'⟩, + refine ⟨c • f', h_f'_meas.const_smul c, _⟩, + exact eventually_eq.fun_comp hff' (λ x, c • x), +end + +lemma const_inner {𝕜 β} [is_R_or_C 𝕜] [normed_add_comm_group β] [inner_product_space 𝕜 β] + {f : α → β} (hfm : ae_strongly_measurable' m f μ) (c : β) : + ae_strongly_measurable' m (λ x, (inner c (f x) : 𝕜)) μ := +begin + rcases hfm with ⟨f', hf'_meas, hf_ae⟩, + refine ⟨λ x, (inner c (f' x) : 𝕜), (@strongly_measurable_const _ _ m _ _).inner hf'_meas, + hf_ae.mono (λ x hx, _)⟩, + dsimp only, + rw hx, +end + +/-- An `m`-strongly measurable function almost everywhere equal to `f`. -/ +noncomputable +def mk (f : α → β) (hfm : ae_strongly_measurable' m f μ) : α → β := hfm.some + +lemma strongly_measurable_mk {f : α → β} (hfm : ae_strongly_measurable' m f μ) : + strongly_measurable[m] (hfm.mk f) := +hfm.some_spec.1 + +lemma ae_eq_mk {f : α → β} (hfm : ae_strongly_measurable' m f μ) : f =ᵐ[μ] hfm.mk f := +hfm.some_spec.2 + +lemma continuous_comp {γ} [topological_space γ] {f : α → β} {g : β → γ} + (hg : continuous g) (hf : ae_strongly_measurable' m f μ) : + ae_strongly_measurable' m (g ∘ f) μ := +⟨λ x, g (hf.mk _ x), + @continuous.comp_strongly_measurable _ _ _ m _ _ _ _ hg hf.strongly_measurable_mk, + hf.ae_eq_mk.mono (λ x hx, by rw [function.comp_apply, hx])⟩ + +end ae_strongly_measurable' + +lemma ae_strongly_measurable'_of_ae_strongly_measurable'_trim {α β} {m m0 m0' : measurable_space α} + [topological_space β] (hm0 : m0 ≤ m0') {μ : measure α} {f : α → β} + (hf : ae_strongly_measurable' m f (μ.trim hm0)) : + ae_strongly_measurable' m f μ := +by { obtain ⟨g, hg_meas, hfg⟩ := hf, exact ⟨g, hg_meas, ae_eq_of_ae_eq_trim hfg⟩, } + +lemma strongly_measurable.ae_strongly_measurable' + {α β} {m m0 : measurable_space α} [topological_space β] + {μ : measure α} {f : α → β} (hf : strongly_measurable[m] f) : + ae_strongly_measurable' m f μ := +⟨f, hf, ae_eq_refl _⟩ + +lemma ae_eq_trim_iff_of_ae_strongly_measurable' {α β} [topological_space β] [metrizable_space β] + {m m0 : measurable_space α} {μ : measure α} {f g : α → β} + (hm : m ≤ m0) (hfm : ae_strongly_measurable' m f μ) (hgm : ae_strongly_measurable' m g μ) : + hfm.mk f =ᵐ[μ.trim hm] hgm.mk g ↔ f =ᵐ[μ] g := +(ae_eq_trim_iff hm hfm.strongly_measurable_mk hgm.strongly_measurable_mk).trans +⟨λ h, hfm.ae_eq_mk.trans (h.trans hgm.ae_eq_mk.symm), + λ h, hfm.ae_eq_mk.symm.trans (h.trans hgm.ae_eq_mk)⟩ + +lemma ae_strongly_measurable.comp_ae_measurable' + {α β γ : Type*} [topological_space β] {mα : measurable_space α} {mγ : measurable_space γ} + {f : α → β} {μ : measure γ} {g : γ → α} + (hf : ae_strongly_measurable f (μ.map g)) (hg : ae_measurable g μ) : + ae_strongly_measurable' (mα.comap g) (f ∘ g) μ := +⟨(hf.mk f) ∘ g, hf.strongly_measurable_mk.comp_measurable (measurable_iff_comap_le.mpr le_rfl), + ae_eq_comp hg hf.ae_eq_mk⟩ + +/-- If the restriction to a set `s` of a σ-algebra `m` is included in the restriction to `s` of +another σ-algebra `m₂` (hypothesis `hs`), the set `s` is `m` measurable and a function `f` almost +everywhere supported on `s` is `m`-ae-strongly-measurable, then `f` is also +`m₂`-ae-strongly-measurable. -/ +lemma ae_strongly_measurable'.ae_strongly_measurable'_of_measurable_space_le_on + {α E} {m m₂ m0 : measurable_space α} {μ : measure α} + [topological_space E] [has_zero E] (hm : m ≤ m0) {s : set α} {f : α → E} + (hs_m : measurable_set[m] s) (hs : ∀ t, measurable_set[m] (s ∩ t) → measurable_set[m₂] (s ∩ t)) + (hf : ae_strongly_measurable' m f μ) (hf_zero : f =ᵐ[μ.restrict sᶜ] 0) : + ae_strongly_measurable' m₂ f μ := +begin + let f' := hf.mk f, + have h_ind_eq : s.indicator (hf.mk f) =ᵐ[μ] f, + { refine filter.eventually_eq.trans _ + (indicator_ae_eq_of_restrict_compl_ae_eq_zero (hm _ hs_m) hf_zero), + filter_upwards [hf.ae_eq_mk] with x hx, + by_cases hxs : x ∈ s, + { simp [hxs, hx], }, + { simp [hxs], }, }, + suffices : strongly_measurable[m₂] (s.indicator (hf.mk f)), + from ae_strongly_measurable'.congr this.ae_strongly_measurable' h_ind_eq, + have hf_ind : strongly_measurable[m] (s.indicator (hf.mk f)), + from hf.strongly_measurable_mk.indicator hs_m, + exact hf_ind.strongly_measurable_of_measurable_space_le_on hs_m hs + (λ x hxs, set.indicator_of_not_mem hxs _), +end + +variables {α E' F F' 𝕜 : Type*} {p : ℝ≥0∞} + [is_R_or_C 𝕜] -- 𝕜 for ℝ or ℂ + -- E' for an inner product space on which we compute integrals + [normed_add_comm_group E'] [inner_product_space 𝕜 E'] + [complete_space E'] [normed_space ℝ E'] + -- F for a Lp submodule + [normed_add_comm_group F] [normed_space 𝕜 F] + -- F' for integrals on a Lp submodule + [normed_add_comm_group F'] [normed_space 𝕜 F'] [normed_space ℝ F'] [complete_space F'] + +section Lp_meas + +/-! ## The subset `Lp_meas` of `Lp` functions a.e. measurable with respect to a sub-sigma-algebra -/ + +variables (F) + +/-- `Lp_meas_subgroup F m p μ` is the subspace of `Lp F p μ` containing functions `f` verifying +`ae_strongly_measurable' m f μ`, i.e. functions which are `μ`-a.e. equal to +an `m`-strongly measurable function. -/ +def Lp_meas_subgroup (m : measurable_space α) [measurable_space α] (p : ℝ≥0∞) (μ : measure α) : + add_subgroup (Lp F p μ) := +{ carrier := {f : (Lp F p μ) | ae_strongly_measurable' m f μ} , + zero_mem' := ⟨(0 : α → F), @strongly_measurable_zero _ _ m _ _, Lp.coe_fn_zero _ _ _⟩, + add_mem' := λ f g hf hg, (hf.add hg).congr (Lp.coe_fn_add f g).symm, + neg_mem' := λ f hf, ae_strongly_measurable'.congr hf.neg (Lp.coe_fn_neg f).symm, } + +variables (𝕜) +/-- `Lp_meas F 𝕜 m p μ` is the subspace of `Lp F p μ` containing functions `f` verifying +`ae_strongly_measurable' m f μ`, i.e. functions which are `μ`-a.e. equal to +an `m`-strongly measurable function. -/ +def Lp_meas (m : measurable_space α) [measurable_space α] (p : ℝ≥0∞) + (μ : measure α) : + submodule 𝕜 (Lp F p μ) := +{ carrier := {f : (Lp F p μ) | ae_strongly_measurable' m f μ} , + zero_mem' := ⟨(0 : α → F), @strongly_measurable_zero _ _ m _ _, Lp.coe_fn_zero _ _ _⟩, + add_mem' := λ f g hf hg, (hf.add hg).congr (Lp.coe_fn_add f g).symm, + smul_mem' := λ c f hf, (hf.const_smul c).congr (Lp.coe_fn_smul c f).symm, } +variables {F 𝕜} + +variables + +lemma mem_Lp_meas_subgroup_iff_ae_strongly_measurable' {m m0 : measurable_space α} {μ : measure α} + {f : Lp F p μ} : + f ∈ Lp_meas_subgroup F m p μ ↔ ae_strongly_measurable' m f μ := +by rw [← add_subgroup.mem_carrier, Lp_meas_subgroup, set.mem_set_of_eq] + +lemma mem_Lp_meas_iff_ae_strongly_measurable' + {m m0 : measurable_space α} {μ : measure α} {f : Lp F p μ} : + f ∈ Lp_meas F 𝕜 m p μ ↔ ae_strongly_measurable' m f μ := +by rw [← set_like.mem_coe, ← submodule.mem_carrier, Lp_meas, set.mem_set_of_eq] + +lemma Lp_meas.ae_strongly_measurable' + {m m0 : measurable_space α} {μ : measure α} (f : Lp_meas F 𝕜 m p μ) : + ae_strongly_measurable' m f μ := +mem_Lp_meas_iff_ae_strongly_measurable'.mp f.mem + +lemma mem_Lp_meas_self + {m0 : measurable_space α} (μ : measure α) (f : Lp F p μ) : + f ∈ Lp_meas F 𝕜 m0 p μ := +mem_Lp_meas_iff_ae_strongly_measurable'.mpr (Lp.ae_strongly_measurable f) + +lemma Lp_meas_subgroup_coe {m m0 : measurable_space α} {μ : measure α} + {f : Lp_meas_subgroup F m p μ} : + ⇑f = (f : Lp F p μ) := +coe_fn_coe_base f + +lemma Lp_meas_coe {m m0 : measurable_space α} {μ : measure α} {f : Lp_meas F 𝕜 m p μ} : + ⇑f = (f : Lp F p μ) := +coe_fn_coe_base f + +lemma mem_Lp_meas_indicator_const_Lp {m m0 : measurable_space α} (hm : m ≤ m0) + {μ : measure α} {s : set α} (hs : measurable_set[m] s) (hμs : μ s ≠ ∞) {c : F} : + indicator_const_Lp p (hm s hs) hμs c ∈ Lp_meas F 𝕜 m p μ := +⟨s.indicator (λ x : α, c), (@strongly_measurable_const _ _ m _ _).indicator hs, + indicator_const_Lp_coe_fn⟩ + +section complete_subspace + +/-! ## The subspace `Lp_meas` is complete. + +We define an `isometry_equiv` between `Lp_meas_subgroup` and the `Lp` space corresponding to the +measure `μ.trim hm`. As a consequence, the completeness of `Lp` implies completeness of +`Lp_meas_subgroup` (and `Lp_meas`). -/ + +variables {ι : Type*} {m m0 : measurable_space α} {μ : measure α} + +/-- If `f` belongs to `Lp_meas_subgroup F m p μ`, then the measurable function it is almost +everywhere equal to (given by `ae_measurable.mk`) belongs to `ℒp` for the measure `μ.trim hm`. -/ +lemma mem_ℒp_trim_of_mem_Lp_meas_subgroup (hm : m ≤ m0) (f : Lp F p μ) + (hf_meas : f ∈ Lp_meas_subgroup F m p μ) : + mem_ℒp (mem_Lp_meas_subgroup_iff_ae_strongly_measurable'.mp hf_meas).some p (μ.trim hm) := +begin + have hf : ae_strongly_measurable' m f μ, + from (mem_Lp_meas_subgroup_iff_ae_strongly_measurable'.mp hf_meas), + let g := hf.some, + obtain ⟨hg, hfg⟩ := hf.some_spec, + change mem_ℒp g p (μ.trim hm), + refine ⟨hg.ae_strongly_measurable, _⟩, + have h_snorm_fg : snorm g p (μ.trim hm) = snorm f p μ, + by { rw snorm_trim hm hg, exact snorm_congr_ae hfg.symm, }, + rw h_snorm_fg, + exact Lp.snorm_lt_top f, +end + +/-- If `f` belongs to `Lp` for the measure `μ.trim hm`, then it belongs to the subgroup +`Lp_meas_subgroup F m p μ`. -/ +lemma mem_Lp_meas_subgroup_to_Lp_of_trim (hm : m ≤ m0) (f : Lp F p (μ.trim hm)) : + (mem_ℒp_of_mem_ℒp_trim hm (Lp.mem_ℒp f)).to_Lp f ∈ Lp_meas_subgroup F m p μ := +begin + let hf_mem_ℒp := mem_ℒp_of_mem_ℒp_trim hm (Lp.mem_ℒp f), + rw mem_Lp_meas_subgroup_iff_ae_strongly_measurable', + refine ae_strongly_measurable'.congr _ (mem_ℒp.coe_fn_to_Lp hf_mem_ℒp).symm, + refine ae_strongly_measurable'_of_ae_strongly_measurable'_trim hm _, + exact Lp.ae_strongly_measurable f, +end + +variables (F p μ) +/-- Map from `Lp_meas_subgroup` to `Lp F p (μ.trim hm)`. -/ +noncomputable +def Lp_meas_subgroup_to_Lp_trim (hm : m ≤ m0) (f : Lp_meas_subgroup F m p μ) : Lp F p (μ.trim hm) := +mem_ℒp.to_Lp (mem_Lp_meas_subgroup_iff_ae_strongly_measurable'.mp f.mem).some + (mem_ℒp_trim_of_mem_Lp_meas_subgroup hm f f.mem) + +variables (𝕜) +/-- Map from `Lp_meas` to `Lp F p (μ.trim hm)`. -/ +noncomputable +def Lp_meas_to_Lp_trim (hm : m ≤ m0) (f : Lp_meas F 𝕜 m p μ) : Lp F p (μ.trim hm) := +mem_ℒp.to_Lp (mem_Lp_meas_iff_ae_strongly_measurable'.mp f.mem).some + (mem_ℒp_trim_of_mem_Lp_meas_subgroup hm f f.mem) +variables {𝕜} + +/-- Map from `Lp F p (μ.trim hm)` to `Lp_meas_subgroup`, inverse of +`Lp_meas_subgroup_to_Lp_trim`. -/ +noncomputable +def Lp_trim_to_Lp_meas_subgroup (hm : m ≤ m0) (f : Lp F p (μ.trim hm)) : Lp_meas_subgroup F m p μ := +⟨(mem_ℒp_of_mem_ℒp_trim hm (Lp.mem_ℒp f)).to_Lp f, mem_Lp_meas_subgroup_to_Lp_of_trim hm f⟩ + +variables (𝕜) +/-- Map from `Lp F p (μ.trim hm)` to `Lp_meas`, inverse of `Lp_meas_to_Lp_trim`. -/ +noncomputable +def Lp_trim_to_Lp_meas (hm : m ≤ m0) (f : Lp F p (μ.trim hm)) : Lp_meas F 𝕜 m p μ := +⟨(mem_ℒp_of_mem_ℒp_trim hm (Lp.mem_ℒp f)).to_Lp f, mem_Lp_meas_subgroup_to_Lp_of_trim hm f⟩ + +variables {F 𝕜 p μ} + +lemma Lp_meas_subgroup_to_Lp_trim_ae_eq (hm : m ≤ m0) (f : Lp_meas_subgroup F m p μ) : + Lp_meas_subgroup_to_Lp_trim F p μ hm f =ᵐ[μ] f := +(ae_eq_of_ae_eq_trim (mem_ℒp.coe_fn_to_Lp (mem_ℒp_trim_of_mem_Lp_meas_subgroup hm ↑f f.mem))).trans + (mem_Lp_meas_subgroup_iff_ae_strongly_measurable'.mp f.mem).some_spec.2.symm + +lemma Lp_trim_to_Lp_meas_subgroup_ae_eq (hm : m ≤ m0) (f : Lp F p (μ.trim hm)) : + Lp_trim_to_Lp_meas_subgroup F p μ hm f =ᵐ[μ] f := +mem_ℒp.coe_fn_to_Lp _ + +lemma Lp_meas_to_Lp_trim_ae_eq (hm : m ≤ m0) (f : Lp_meas F 𝕜 m p μ) : + Lp_meas_to_Lp_trim F 𝕜 p μ hm f =ᵐ[μ] f := +(ae_eq_of_ae_eq_trim (mem_ℒp.coe_fn_to_Lp (mem_ℒp_trim_of_mem_Lp_meas_subgroup hm ↑f f.mem))).trans + (mem_Lp_meas_subgroup_iff_ae_strongly_measurable'.mp f.mem).some_spec.2.symm + +lemma Lp_trim_to_Lp_meas_ae_eq (hm : m ≤ m0) (f : Lp F p (μ.trim hm)) : + Lp_trim_to_Lp_meas F 𝕜 p μ hm f =ᵐ[μ] f := +mem_ℒp.coe_fn_to_Lp _ + +/-- `Lp_trim_to_Lp_meas_subgroup` is a right inverse of `Lp_meas_subgroup_to_Lp_trim`. -/ +lemma Lp_meas_subgroup_to_Lp_trim_right_inv (hm : m ≤ m0) : + function.right_inverse (Lp_trim_to_Lp_meas_subgroup F p μ hm) + (Lp_meas_subgroup_to_Lp_trim F p μ hm) := +begin + intro f, + ext1, + refine ae_eq_trim_of_strongly_measurable hm + (Lp.strongly_measurable _) (Lp.strongly_measurable _) _, + exact (Lp_meas_subgroup_to_Lp_trim_ae_eq hm _).trans (Lp_trim_to_Lp_meas_subgroup_ae_eq hm _), +end + +/-- `Lp_trim_to_Lp_meas_subgroup` is a left inverse of `Lp_meas_subgroup_to_Lp_trim`. -/ +lemma Lp_meas_subgroup_to_Lp_trim_left_inv (hm : m ≤ m0) : + function.left_inverse (Lp_trim_to_Lp_meas_subgroup F p μ hm) + (Lp_meas_subgroup_to_Lp_trim F p μ hm) := +begin + intro f, + ext1, + ext1, + rw ← Lp_meas_subgroup_coe, + exact (Lp_trim_to_Lp_meas_subgroup_ae_eq hm _).trans (Lp_meas_subgroup_to_Lp_trim_ae_eq hm _), +end + +lemma Lp_meas_subgroup_to_Lp_trim_add (hm : m ≤ m0) (f g : Lp_meas_subgroup F m p μ) : + Lp_meas_subgroup_to_Lp_trim F p μ hm (f + g) + = Lp_meas_subgroup_to_Lp_trim F p μ hm f + Lp_meas_subgroup_to_Lp_trim F p μ hm g := +begin + ext1, + refine eventually_eq.trans _ (Lp.coe_fn_add _ _).symm, + refine ae_eq_trim_of_strongly_measurable hm (Lp.strongly_measurable _) _ _, + { exact (Lp.strongly_measurable _).add (Lp.strongly_measurable _), }, + refine (Lp_meas_subgroup_to_Lp_trim_ae_eq hm _).trans _, + refine eventually_eq.trans _ + (eventually_eq.add (Lp_meas_subgroup_to_Lp_trim_ae_eq hm f).symm + (Lp_meas_subgroup_to_Lp_trim_ae_eq hm g).symm), + refine (Lp.coe_fn_add _ _).trans _, + simp_rw Lp_meas_subgroup_coe, + exact eventually_of_forall (λ x, by refl), +end + +lemma Lp_meas_subgroup_to_Lp_trim_neg (hm : m ≤ m0) (f : Lp_meas_subgroup F m p μ) : + Lp_meas_subgroup_to_Lp_trim F p μ hm (-f) + = -Lp_meas_subgroup_to_Lp_trim F p μ hm f := +begin + ext1, + refine eventually_eq.trans _ (Lp.coe_fn_neg _).symm, + refine ae_eq_trim_of_strongly_measurable hm (Lp.strongly_measurable _) _ _, + { exact @strongly_measurable.neg _ _ _ m _ _ _ (Lp.strongly_measurable _), }, + refine (Lp_meas_subgroup_to_Lp_trim_ae_eq hm _).trans _, + refine eventually_eq.trans _ + (eventually_eq.neg (Lp_meas_subgroup_to_Lp_trim_ae_eq hm f).symm), + refine (Lp.coe_fn_neg _).trans _, + simp_rw Lp_meas_subgroup_coe, + exact eventually_of_forall (λ x, by refl), +end + +lemma Lp_meas_subgroup_to_Lp_trim_sub (hm : m ≤ m0) (f g : Lp_meas_subgroup F m p μ) : + Lp_meas_subgroup_to_Lp_trim F p μ hm (f - g) + = Lp_meas_subgroup_to_Lp_trim F p μ hm f - Lp_meas_subgroup_to_Lp_trim F p μ hm g := +by rw [sub_eq_add_neg, sub_eq_add_neg, Lp_meas_subgroup_to_Lp_trim_add, + Lp_meas_subgroup_to_Lp_trim_neg] + +lemma Lp_meas_to_Lp_trim_smul (hm : m ≤ m0) (c : 𝕜) (f : Lp_meas F 𝕜 m p μ) : + Lp_meas_to_Lp_trim F 𝕜 p μ hm (c • f) = c • Lp_meas_to_Lp_trim F 𝕜 p μ hm f := +begin + ext1, + refine eventually_eq.trans _ (Lp.coe_fn_smul _ _).symm, + refine ae_eq_trim_of_strongly_measurable hm (Lp.strongly_measurable _) _ _, + { exact (Lp.strongly_measurable _).const_smul c, }, + refine (Lp_meas_to_Lp_trim_ae_eq hm _).trans _, + refine (Lp.coe_fn_smul _ _).trans _, + refine (Lp_meas_to_Lp_trim_ae_eq hm f).mono (λ x hx, _), + rw [pi.smul_apply, pi.smul_apply, hx], + refl, +end + +/-- `Lp_meas_subgroup_to_Lp_trim` preserves the norm. -/ +lemma Lp_meas_subgroup_to_Lp_trim_norm_map [hp : fact (1 ≤ p)] (hm : m ≤ m0) + (f : Lp_meas_subgroup F m p μ) : + ‖Lp_meas_subgroup_to_Lp_trim F p μ hm f‖ = ‖f‖ := +begin + rw [Lp.norm_def, snorm_trim hm (Lp.strongly_measurable _), + snorm_congr_ae (Lp_meas_subgroup_to_Lp_trim_ae_eq hm _), Lp_meas_subgroup_coe, ← Lp.norm_def], + congr, +end + +lemma isometry_Lp_meas_subgroup_to_Lp_trim [hp : fact (1 ≤ p)] (hm : m ≤ m0) : + isometry (Lp_meas_subgroup_to_Lp_trim F p μ hm) := +isometry.of_dist_eq $ λ f g, by rw [dist_eq_norm, ← Lp_meas_subgroup_to_Lp_trim_sub, + Lp_meas_subgroup_to_Lp_trim_norm_map, dist_eq_norm] + +variables (F p μ) +/-- `Lp_meas_subgroup` and `Lp F p (μ.trim hm)` are isometric. -/ +noncomputable +def Lp_meas_subgroup_to_Lp_trim_iso [hp : fact (1 ≤ p)] (hm : m ≤ m0) : + Lp_meas_subgroup F m p μ ≃ᵢ Lp F p (μ.trim hm) := +{ to_fun := Lp_meas_subgroup_to_Lp_trim F p μ hm, + inv_fun := Lp_trim_to_Lp_meas_subgroup F p μ hm, + left_inv := Lp_meas_subgroup_to_Lp_trim_left_inv hm, + right_inv := Lp_meas_subgroup_to_Lp_trim_right_inv hm, + isometry_to_fun := isometry_Lp_meas_subgroup_to_Lp_trim hm, } + +variables (𝕜) +/-- `Lp_meas_subgroup` and `Lp_meas` are isometric. -/ +noncomputable +def Lp_meas_subgroup_to_Lp_meas_iso [hp : fact (1 ≤ p)] : + Lp_meas_subgroup F m p μ ≃ᵢ Lp_meas F 𝕜 m p μ := +isometry_equiv.refl (Lp_meas_subgroup F m p μ) + +/-- `Lp_meas` and `Lp F p (μ.trim hm)` are isometric, with a linear equivalence. -/ +noncomputable +def Lp_meas_to_Lp_trim_lie [hp : fact (1 ≤ p)] (hm : m ≤ m0) : + Lp_meas F 𝕜 m p μ ≃ₗᵢ[𝕜] Lp F p (μ.trim hm) := +{ to_fun := Lp_meas_to_Lp_trim F 𝕜 p μ hm, + inv_fun := Lp_trim_to_Lp_meas F 𝕜 p μ hm, + left_inv := Lp_meas_subgroup_to_Lp_trim_left_inv hm, + right_inv := Lp_meas_subgroup_to_Lp_trim_right_inv hm, + map_add' := Lp_meas_subgroup_to_Lp_trim_add hm, + map_smul' := Lp_meas_to_Lp_trim_smul hm, + norm_map' := Lp_meas_subgroup_to_Lp_trim_norm_map hm, } +variables {F 𝕜 p μ} + +instance [hm : fact (m ≤ m0)] [complete_space F] [hp : fact (1 ≤ p)] : + complete_space (Lp_meas_subgroup F m p μ) := +by { rw (Lp_meas_subgroup_to_Lp_trim_iso F p μ hm.elim).complete_space_iff, apply_instance, } + +-- For now just no-lint this; lean4's tree-based logging will make this easier to debug. +-- One possible change might be to generalize `𝕜` from `is_R_or_C` to `normed_field`, as this +-- result may well hold there. +@[nolint fails_quickly] +instance [hm : fact (m ≤ m0)] [complete_space F] [hp : fact (1 ≤ p)] : + complete_space (Lp_meas F 𝕜 m p μ) := +by { rw (Lp_meas_subgroup_to_Lp_meas_iso F 𝕜 p μ).symm.complete_space_iff, apply_instance, } + +lemma is_complete_ae_strongly_measurable' [hp : fact (1 ≤ p)] [complete_space F] (hm : m ≤ m0) : + is_complete {f : Lp F p μ | ae_strongly_measurable' m f μ} := +begin + rw ← complete_space_coe_iff_is_complete, + haveI : fact (m ≤ m0) := ⟨hm⟩, + change complete_space (Lp_meas_subgroup F m p μ), + apply_instance, +end + +lemma is_closed_ae_strongly_measurable' [hp : fact (1 ≤ p)] [complete_space F] (hm : m ≤ m0) : + is_closed {f : Lp F p μ | ae_strongly_measurable' m f μ} := +is_complete.is_closed (is_complete_ae_strongly_measurable' hm) + +end complete_subspace + +section strongly_measurable + +variables {m m0 : measurable_space α} {μ : measure α} + +/-- We do not get `ae_fin_strongly_measurable f (μ.trim hm)`, since we don't have +`f =ᵐ[μ.trim hm] Lp_meas_to_Lp_trim F 𝕜 p μ hm f` but only the weaker +`f =ᵐ[μ] Lp_meas_to_Lp_trim F 𝕜 p μ hm f`. -/ +lemma Lp_meas.ae_fin_strongly_measurable' (hm : m ≤ m0) (f : Lp_meas F 𝕜 m p μ) (hp_ne_zero : p ≠ 0) + (hp_ne_top : p ≠ ∞) : + ∃ g, fin_strongly_measurable g (μ.trim hm) ∧ f =ᵐ[μ] g := +⟨Lp_meas_subgroup_to_Lp_trim F p μ hm f, Lp.fin_strongly_measurable _ hp_ne_zero hp_ne_top, + (Lp_meas_subgroup_to_Lp_trim_ae_eq hm f).symm⟩ + +/-- When applying the inverse of `Lp_meas_to_Lp_trim_lie` (which takes a function in the Lp space of +the sub-sigma algebra and returns its version in the larger Lp space) to an indicator of the +sub-sigma-algebra, we obtain an indicator in the Lp space of the larger sigma-algebra. -/ +lemma Lp_meas_to_Lp_trim_lie_symm_indicator [one_le_p : fact (1 ≤ p)] [normed_space ℝ F] + {hm : m ≤ m0} {s : set α} {μ : measure α} + (hs : measurable_set[m] s) (hμs : μ.trim hm s ≠ ∞) (c : F) : + ((Lp_meas_to_Lp_trim_lie F ℝ p μ hm).symm + (indicator_const_Lp p hs hμs c) : Lp F p μ) + = indicator_const_Lp p (hm s hs) ((le_trim hm).trans_lt hμs.lt_top).ne c := +begin + ext1, + rw ← Lp_meas_coe, + change Lp_trim_to_Lp_meas F ℝ p μ hm (indicator_const_Lp p hs hμs c) + =ᵐ[μ] (indicator_const_Lp p _ _ c : α → F), + refine (Lp_trim_to_Lp_meas_ae_eq hm _).trans _, + exact (ae_eq_of_ae_eq_trim indicator_const_Lp_coe_fn).trans indicator_const_Lp_coe_fn.symm, +end + +lemma Lp_meas_to_Lp_trim_lie_symm_to_Lp [one_le_p : fact (1 ≤ p)] [normed_space ℝ F] + (hm : m ≤ m0) (f : α → F) (hf : mem_ℒp f p (μ.trim hm)) : + ((Lp_meas_to_Lp_trim_lie F ℝ p μ hm).symm (hf.to_Lp f) : Lp F p μ) + = (mem_ℒp_of_mem_ℒp_trim hm hf).to_Lp f := +begin + ext1, + rw ← Lp_meas_coe, + refine (Lp_trim_to_Lp_meas_ae_eq hm _).trans _, + exact (ae_eq_of_ae_eq_trim (mem_ℒp.coe_fn_to_Lp hf)).trans (mem_ℒp.coe_fn_to_Lp _).symm, +end + +end strongly_measurable + +end Lp_meas + + +section induction + +variables {m m0 : measurable_space α} {μ : measure α} [fact (1 ≤ p)] [normed_space ℝ F] + +/-- Auxiliary lemma for `Lp.induction_strongly_measurable`. -/ +@[elab_as_eliminator] +lemma Lp.induction_strongly_measurable_aux (hm : m ≤ m0) (hp_ne_top : p ≠ ∞) (P : Lp F p μ → Prop) + (h_ind : ∀ (c : F) {s : set α} (hs : measurable_set[m] s) (hμs : μ s < ∞), + P (Lp.simple_func.indicator_const p (hm s hs) hμs.ne c)) + (h_add : ∀ ⦃f g⦄, ∀ hf : mem_ℒp f p μ, ∀ hg : mem_ℒp g p μ, + ∀ hfm : ae_strongly_measurable' m f μ, ∀ hgm : ae_strongly_measurable' m g μ, + disjoint (function.support f) (function.support g) → + P (hf.to_Lp f) → P (hg.to_Lp g) → P ((hf.to_Lp f) + (hg.to_Lp g))) + (h_closed : is_closed {f : Lp_meas F ℝ m p μ | P f}) : + ∀ f : Lp F p μ, ae_strongly_measurable' m f μ → P f := +begin + intros f hf, + let f' := (⟨f, hf⟩ : Lp_meas F ℝ m p μ), + let g := Lp_meas_to_Lp_trim_lie F ℝ p μ hm f', + have hfg : f' = (Lp_meas_to_Lp_trim_lie F ℝ p μ hm).symm g, + by simp only [linear_isometry_equiv.symm_apply_apply], + change P ↑f', + rw hfg, + refine @Lp.induction α F m _ p (μ.trim hm) _ hp_ne_top + (λ g, P ((Lp_meas_to_Lp_trim_lie F ℝ p μ hm).symm g)) _ _ _ g, + { intros b t ht hμt, + rw [Lp.simple_func.coe_indicator_const, + Lp_meas_to_Lp_trim_lie_symm_indicator ht hμt.ne b], + have hμt' : μ t < ∞, from (le_trim hm).trans_lt hμt, + specialize h_ind b ht hμt', + rwa Lp.simple_func.coe_indicator_const at h_ind, }, + { intros f g hf hg h_disj hfP hgP, + rw linear_isometry_equiv.map_add, + push_cast, + have h_eq : ∀ (f : α → F) (hf : mem_ℒp f p (μ.trim hm)), + ((Lp_meas_to_Lp_trim_lie F ℝ p μ hm).symm (mem_ℒp.to_Lp f hf) : Lp F p μ) + = (mem_ℒp_of_mem_ℒp_trim hm hf).to_Lp f, + from Lp_meas_to_Lp_trim_lie_symm_to_Lp hm, + rw h_eq f hf at hfP ⊢, + rw h_eq g hg at hgP ⊢, + exact h_add (mem_ℒp_of_mem_ℒp_trim hm hf) (mem_ℒp_of_mem_ℒp_trim hm hg) + (ae_strongly_measurable'_of_ae_strongly_measurable'_trim hm hf.ae_strongly_measurable) + (ae_strongly_measurable'_of_ae_strongly_measurable'_trim hm hg.ae_strongly_measurable) + h_disj hfP hgP, }, + { change is_closed ((Lp_meas_to_Lp_trim_lie F ℝ p μ hm).symm ⁻¹' {g : Lp_meas F ℝ m p μ | P ↑g}), + exact is_closed.preimage (linear_isometry_equiv.continuous _) h_closed, }, +end + +/-- To prove something for an `Lp` function a.e. strongly measurable with respect to a +sub-σ-algebra `m` in a normed space, it suffices to show that +* the property holds for (multiples of) characteristic functions which are measurable w.r.t. `m`; +* is closed under addition; +* the set of functions in `Lp` strongly measurable w.r.t. `m` for which the property holds is + closed. +-/ +@[elab_as_eliminator] +lemma Lp.induction_strongly_measurable (hm : m ≤ m0) (hp_ne_top : p ≠ ∞) (P : Lp F p μ → Prop) + (h_ind : ∀ (c : F) {s : set α} (hs : measurable_set[m] s) (hμs : μ s < ∞), + P (Lp.simple_func.indicator_const p (hm s hs) hμs.ne c)) + (h_add : ∀ ⦃f g⦄, ∀ hf : mem_ℒp f p μ, ∀ hg : mem_ℒp g p μ, + ∀ hfm : strongly_measurable[m] f, ∀ hgm : strongly_measurable[m] g, + disjoint (function.support f) (function.support g) → + P (hf.to_Lp f) → P (hg.to_Lp g) → P ((hf.to_Lp f) + (hg.to_Lp g))) + (h_closed : is_closed {f : Lp_meas F ℝ m p μ | P f}) : + ∀ f : Lp F p μ, ae_strongly_measurable' m f μ → P f := +begin + intros f hf, + suffices h_add_ae : ∀ ⦃f g⦄, ∀ hf : mem_ℒp f p μ, ∀ hg : mem_ℒp g p μ, + ∀ hfm : ae_strongly_measurable' m f μ, ∀ hgm : ae_strongly_measurable' m g μ, + disjoint (function.support f) (function.support g) → + P (hf.to_Lp f) → P (hg.to_Lp g) → P ((hf.to_Lp f) + (hg.to_Lp g)), + from Lp.induction_strongly_measurable_aux hm hp_ne_top P h_ind h_add_ae h_closed f hf, + intros f g hf hg hfm hgm h_disj hPf hPg, + let s_f : set α := function.support (hfm.mk f), + have hs_f : measurable_set[m] s_f := hfm.strongly_measurable_mk.measurable_set_support, + have hs_f_eq : s_f =ᵐ[μ] function.support f := hfm.ae_eq_mk.symm.support, + let s_g : set α := function.support (hgm.mk g), + have hs_g : measurable_set[m] s_g := hgm.strongly_measurable_mk.measurable_set_support, + have hs_g_eq : s_g =ᵐ[μ] function.support g := hgm.ae_eq_mk.symm.support, + have h_inter_empty : ((s_f ∩ s_g) : set α) =ᵐ[μ] (∅ : set α), + { refine (hs_f_eq.inter hs_g_eq).trans _, + suffices : function.support f ∩ function.support g = ∅, by rw this, + exact set.disjoint_iff_inter_eq_empty.mp h_disj, }, + let f' := (s_f \ s_g).indicator (hfm.mk f), + have hff' : f =ᵐ[μ] f', + { have : s_f \ s_g =ᵐ[μ] s_f, + { rw [← set.diff_inter_self_eq_diff, set.inter_comm], + refine ((ae_eq_refl s_f).diff h_inter_empty).trans _, + rw set.diff_empty, }, + refine ((indicator_ae_eq_of_ae_eq_set this).trans _).symm, + rw set.indicator_support, + exact hfm.ae_eq_mk.symm, }, + have hf'_meas : strongly_measurable[m] f', + from hfm.strongly_measurable_mk.indicator (hs_f.diff hs_g), + have hf'_Lp : mem_ℒp f' p μ := hf.ae_eq hff', + let g' := (s_g \ s_f).indicator (hgm.mk g), + have hgg' : g =ᵐ[μ] g', + { have : s_g \ s_f =ᵐ[μ] s_g, + { rw [← set.diff_inter_self_eq_diff], + refine ((ae_eq_refl s_g).diff h_inter_empty).trans _, + rw set.diff_empty, }, + refine ((indicator_ae_eq_of_ae_eq_set this).trans _).symm, + rw set.indicator_support, + exact hgm.ae_eq_mk.symm, }, + have hg'_meas : strongly_measurable[m] g', + from hgm.strongly_measurable_mk.indicator (hs_g.diff hs_f), + have hg'_Lp : mem_ℒp g' p μ := hg.ae_eq hgg', + have h_disj : disjoint (function.support f') (function.support g'), + { have : disjoint (s_f \ s_g) (s_g \ s_f) := disjoint_sdiff_sdiff, + exact this.mono set.support_indicator_subset set.support_indicator_subset, }, + rw ← mem_ℒp.to_Lp_congr hf'_Lp hf hff'.symm at ⊢ hPf, + rw ← mem_ℒp.to_Lp_congr hg'_Lp hg hgg'.symm at ⊢ hPg, + exact h_add hf'_Lp hg'_Lp hf'_meas hg'_meas h_disj hPf hPg, +end + +/-- To prove something for an arbitrary `mem_ℒp` function a.e. strongly measurable with respect +to a sub-σ-algebra `m` in a normed space, it suffices to show that +* the property holds for (multiples of) characteristic functions which are measurable w.r.t. `m`; +* is closed under addition; +* the set of functions in the `Lᵖ` space strongly measurable w.r.t. `m` for which the property + holds is closed. +* the property is closed under the almost-everywhere equal relation. +-/ +@[elab_as_eliminator] +lemma mem_ℒp.induction_strongly_measurable (hm : m ≤ m0) (hp_ne_top : p ≠ ∞) + (P : (α → F) → Prop) + (h_ind : ∀ (c : F) ⦃s⦄, measurable_set[m] s → μ s < ∞ → P (s.indicator (λ _, c))) + (h_add : ∀ ⦃f g : α → F⦄, disjoint (function.support f) (function.support g) + → mem_ℒp f p μ → mem_ℒp g p μ → strongly_measurable[m] f → strongly_measurable[m] g → + P f → P g → P (f + g)) + (h_closed : is_closed {f : Lp_meas F ℝ m p μ | P f} ) + (h_ae : ∀ ⦃f g⦄, f =ᵐ[μ] g → mem_ℒp f p μ → P f → P g) : + ∀ ⦃f : α → F⦄ (hf : mem_ℒp f p μ) (hfm : ae_strongly_measurable' m f μ), P f := +begin + intros f hf hfm, + let f_Lp := hf.to_Lp f, + have hfm_Lp : ae_strongly_measurable' m f_Lp μ, from hfm.congr hf.coe_fn_to_Lp.symm, + refine h_ae (hf.coe_fn_to_Lp) (Lp.mem_ℒp _) _, + change P f_Lp, + refine Lp.induction_strongly_measurable hm hp_ne_top (λ f, P ⇑f) _ _ h_closed f_Lp hfm_Lp, + { intros c s hs hμs, + rw Lp.simple_func.coe_indicator_const, + refine h_ae (indicator_const_Lp_coe_fn).symm _ (h_ind c hs hμs), + exact mem_ℒp_indicator_const p (hm s hs) c (or.inr hμs.ne), }, + { intros f g hf_mem hg_mem hfm hgm h_disj hfP hgP, + have hfP' : P f := h_ae (hf_mem.coe_fn_to_Lp) (Lp.mem_ℒp _) hfP, + have hgP' : P g := h_ae (hg_mem.coe_fn_to_Lp) (Lp.mem_ℒp _) hgP, + specialize h_add h_disj hf_mem hg_mem hfm hgm hfP' hgP', + refine h_ae _ (hf_mem.add hg_mem) h_add, + exact ((hf_mem.coe_fn_to_Lp).symm.add (hg_mem.coe_fn_to_Lp).symm).trans + (Lp.coe_fn_add _ _).symm, }, +end + +end induction + +end measure_theory diff --git a/src/measure_theory/function/conditional_expectation/basic.lean b/src/measure_theory/function/conditional_expectation/basic.lean index 1bd8bbd85c74e..404015d07d653 100644 --- a/src/measure_theory/function/conditional_expectation/basic.lean +++ b/src/measure_theory/function/conditional_expectation/basic.lean @@ -3,10 +3,7 @@ Copyright (c) 2021 Rémy Degenne. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ - -import analysis.inner_product_space.projection -import measure_theory.function.l2_space -import measure_theory.function.ae_eq_of_integral +import measure_theory.function.conditional_expectation.condexp_L1 /-! # Conditional expectation @@ -29,6 +26,10 @@ The construction is done in four steps: `α → E` equal to 0 if `f` is not integrable, and equal to an `m`-measurable representative of `condexp_L1_clm` applied to `[f]`, the equivalence class of `f` in `L¹`. +The first step is done in `measure_theory.function.conditional_expectation.condexp_L2`, the two +next steps in `measure_theory.function.conditional_expectation.condexp_L1` and the final step is +performed in this file. + ## Main results The conditional expectation and its properties @@ -46,11 +47,6 @@ linear map `condexp_L1_clm` from `L1` to `L1`. `condexp` should be used in most Uniqueness of the conditional expectation -* `Lp.ae_eq_of_forall_set_integral_eq'`: two `Lp` functions verifying the equality of integrals - defining the conditional expectation are equal. -* `ae_eq_of_forall_set_integral_eq_of_sigma_finite'`: two functions verifying the equality of - integrals defining the conditional expectation are equal almost everywhere. - Requires `[sigma_finite (μ.trim hm)]`. * `ae_eq_condexp_of_forall_set_integral_eq`: an a.e. `m`-measurable function which verifies the equality of integrals is a.e. equal to `condexp`. @@ -60,1859 +56,23 @@ For a measure `μ` defined on a measurable space structure `m0`, another measura `m` with `hm : m ≤ m0` (a sub-σ-algebra) and a function `f`, we define the notation * `μ[f|m] = condexp m μ f`. -## Implementation notes - -Most of the results in this file are valid for a complete real normed space `F`. -However, some lemmas also use `𝕜 : is_R_or_C`: -* `condexp_L2` is defined only for an `inner_product_space` for now, and we use `𝕜` for its field. -* results about scalar multiplication are stated not only for `ℝ` but also for `𝕜` if we happen to - have `normed_space 𝕜 F`. - ## Tags conditional expectation, conditional expected value -/ -noncomputable theory -open topological_space measure_theory.Lp filter continuous_linear_map -open_locale nnreal ennreal topology big_operators measure_theory +open topological_space measure_theory.Lp filter +open_locale ennreal topology big_operators measure_theory namespace measure_theory -/-- A function `f` verifies `ae_strongly_measurable' m f μ` if it is `μ`-a.e. equal to -an `m`-strongly measurable function. This is similar to `ae_strongly_measurable`, but the -`measurable_space` structures used for the measurability statement and for the measure are -different. -/ -def ae_strongly_measurable' {α β} [topological_space β] - (m : measurable_space α) {m0 : measurable_space α} - (f : α → β) (μ : measure α) : Prop := -∃ g : α → β, strongly_measurable[m] g ∧ f =ᵐ[μ] g - -namespace ae_strongly_measurable' - -variables {α β 𝕜 : Type*} {m m0 : measurable_space α} {μ : measure α} - [topological_space β] {f g : α → β} - -lemma congr (hf : ae_strongly_measurable' m f μ) (hfg : f =ᵐ[μ] g) : - ae_strongly_measurable' m g μ := -by { obtain ⟨f', hf'_meas, hff'⟩ := hf, exact ⟨f', hf'_meas, hfg.symm.trans hff'⟩, } - -lemma add [has_add β] [has_continuous_add β] (hf : ae_strongly_measurable' m f μ) - (hg : ae_strongly_measurable' m g μ) : - ae_strongly_measurable' m (f+g) μ := -begin - rcases hf with ⟨f', h_f'_meas, hff'⟩, - rcases hg with ⟨g', h_g'_meas, hgg'⟩, - exact ⟨f' + g', h_f'_meas.add h_g'_meas, hff'.add hgg'⟩, -end - -lemma neg [add_group β] [topological_add_group β] - {f : α → β} (hfm : ae_strongly_measurable' m f μ) : - ae_strongly_measurable' m (-f) μ := -begin - rcases hfm with ⟨f', hf'_meas, hf_ae⟩, - refine ⟨-f', hf'_meas.neg, hf_ae.mono (λ x hx, _)⟩, - simp_rw pi.neg_apply, - rw hx, -end - -lemma sub [add_group β] [topological_add_group β] {f g : α → β} - (hfm : ae_strongly_measurable' m f μ) (hgm : ae_strongly_measurable' m g μ) : - ae_strongly_measurable' m (f - g) μ := -begin - rcases hfm with ⟨f', hf'_meas, hf_ae⟩, - rcases hgm with ⟨g', hg'_meas, hg_ae⟩, - refine ⟨f'-g', hf'_meas.sub hg'_meas, hf_ae.mp (hg_ae.mono (λ x hx1 hx2, _))⟩, - simp_rw pi.sub_apply, - rw [hx1, hx2], -end - -lemma const_smul [has_smul 𝕜 β] [has_continuous_const_smul 𝕜 β] - (c : 𝕜) (hf : ae_strongly_measurable' m f μ) : - ae_strongly_measurable' m (c • f) μ := -begin - rcases hf with ⟨f', h_f'_meas, hff'⟩, - refine ⟨c • f', h_f'_meas.const_smul c, _⟩, - exact eventually_eq.fun_comp hff' (λ x, c • x), -end - -lemma const_inner {𝕜 β} [is_R_or_C 𝕜] [normed_add_comm_group β] [inner_product_space 𝕜 β] - {f : α → β} (hfm : ae_strongly_measurable' m f μ) (c : β) : - ae_strongly_measurable' m (λ x, (inner c (f x) : 𝕜)) μ := -begin - rcases hfm with ⟨f', hf'_meas, hf_ae⟩, - refine ⟨λ x, (inner c (f' x) : 𝕜), (@strongly_measurable_const _ _ m _ _).inner hf'_meas, - hf_ae.mono (λ x hx, _)⟩, - dsimp only, - rw hx, -end - -/-- An `m`-strongly measurable function almost everywhere equal to `f`. -/ -def mk (f : α → β) (hfm : ae_strongly_measurable' m f μ) : α → β := hfm.some - -lemma strongly_measurable_mk {f : α → β} (hfm : ae_strongly_measurable' m f μ) : - strongly_measurable[m] (hfm.mk f) := -hfm.some_spec.1 - -lemma ae_eq_mk {f : α → β} (hfm : ae_strongly_measurable' m f μ) : f =ᵐ[μ] hfm.mk f := -hfm.some_spec.2 - -lemma continuous_comp {γ} [topological_space γ] {f : α → β} {g : β → γ} - (hg : continuous g) (hf : ae_strongly_measurable' m f μ) : - ae_strongly_measurable' m (g ∘ f) μ := -⟨λ x, g (hf.mk _ x), - @continuous.comp_strongly_measurable _ _ _ m _ _ _ _ hg hf.strongly_measurable_mk, - hf.ae_eq_mk.mono (λ x hx, by rw [function.comp_apply, hx])⟩ - -end ae_strongly_measurable' - -lemma ae_strongly_measurable'_of_ae_strongly_measurable'_trim {α β} {m m0 m0' : measurable_space α} - [topological_space β] (hm0 : m0 ≤ m0') {μ : measure α} {f : α → β} - (hf : ae_strongly_measurable' m f (μ.trim hm0)) : - ae_strongly_measurable' m f μ := -by { obtain ⟨g, hg_meas, hfg⟩ := hf, exact ⟨g, hg_meas, ae_eq_of_ae_eq_trim hfg⟩, } - -lemma strongly_measurable.ae_strongly_measurable' - {α β} {m m0 : measurable_space α} [topological_space β] - {μ : measure α} {f : α → β} (hf : strongly_measurable[m] f) : - ae_strongly_measurable' m f μ := -⟨f, hf, ae_eq_refl _⟩ - -lemma ae_eq_trim_iff_of_ae_strongly_measurable' {α β} [topological_space β] [metrizable_space β] - {m m0 : measurable_space α} {μ : measure α} {f g : α → β} - (hm : m ≤ m0) (hfm : ae_strongly_measurable' m f μ) (hgm : ae_strongly_measurable' m g μ) : - hfm.mk f =ᵐ[μ.trim hm] hgm.mk g ↔ f =ᵐ[μ] g := -(ae_eq_trim_iff hm hfm.strongly_measurable_mk hgm.strongly_measurable_mk).trans -⟨λ h, hfm.ae_eq_mk.trans (h.trans hgm.ae_eq_mk.symm), - λ h, hfm.ae_eq_mk.symm.trans (h.trans hgm.ae_eq_mk)⟩ - -lemma ae_strongly_measurable.comp_ae_measurable' - {α β γ : Type*} [topological_space β] {mα : measurable_space α} {mγ : measurable_space γ} - {f : α → β} {μ : measure γ} {g : γ → α} - (hf : ae_strongly_measurable f (μ.map g)) (hg : ae_measurable g μ) : - ae_strongly_measurable' (mα.comap g) (f ∘ g) μ := -⟨(hf.mk f) ∘ g, hf.strongly_measurable_mk.comp_measurable (measurable_iff_comap_le.mpr le_rfl), - ae_eq_comp hg hf.ae_eq_mk⟩ - -/-- If the restriction to a set `s` of a σ-algebra `m` is included in the restriction to `s` of -another σ-algebra `m₂` (hypothesis `hs`), the set `s` is `m` measurable and a function `f` almost -everywhere supported on `s` is `m`-ae-strongly-measurable, then `f` is also -`m₂`-ae-strongly-measurable. -/ -lemma ae_strongly_measurable'.ae_strongly_measurable'_of_measurable_space_le_on - {α E} {m m₂ m0 : measurable_space α} {μ : measure α} - [topological_space E] [has_zero E] (hm : m ≤ m0) {s : set α} {f : α → E} - (hs_m : measurable_set[m] s) (hs : ∀ t, measurable_set[m] (s ∩ t) → measurable_set[m₂] (s ∩ t)) - (hf : ae_strongly_measurable' m f μ) (hf_zero : f =ᵐ[μ.restrict sᶜ] 0) : - ae_strongly_measurable' m₂ f μ := -begin - let f' := hf.mk f, - have h_ind_eq : s.indicator (hf.mk f) =ᵐ[μ] f, - { refine filter.eventually_eq.trans _ - (indicator_ae_eq_of_restrict_compl_ae_eq_zero (hm _ hs_m) hf_zero), - filter_upwards [hf.ae_eq_mk] with x hx, - by_cases hxs : x ∈ s, - { simp [hxs, hx], }, - { simp [hxs], }, }, - suffices : strongly_measurable[m₂] (s.indicator (hf.mk f)), - from ae_strongly_measurable'.congr this.ae_strongly_measurable' h_ind_eq, - have hf_ind : strongly_measurable[m] (s.indicator (hf.mk f)), - from hf.strongly_measurable_mk.indicator hs_m, - exact hf_ind.strongly_measurable_of_measurable_space_le_on hs_m hs - (λ x hxs, set.indicator_of_not_mem hxs _), -end - -variables {α β γ E E' F F' G G' H 𝕜 : Type*} {p : ℝ≥0∞} +variables {α F F' 𝕜 : Type*} {p : ℝ≥0∞} [is_R_or_C 𝕜] -- 𝕜 for ℝ or ℂ - [topological_space β] -- β for a generic topological space - -- E for an inner product space - [normed_add_comm_group E] [inner_product_space 𝕜 E] - -- E' for an inner product space on which we compute integrals - [normed_add_comm_group E'] [inner_product_space 𝕜 E'] - [complete_space E'] [normed_space ℝ E'] -- F for a Lp submodule [normed_add_comm_group F] [normed_space 𝕜 F] -- F' for integrals on a Lp submodule [normed_add_comm_group F'] [normed_space 𝕜 F'] [normed_space ℝ F'] [complete_space F'] - -- G for a Lp add_subgroup - [normed_add_comm_group G] - -- G' for integrals on a Lp add_subgroup - [normed_add_comm_group G'] [normed_space ℝ G'] [complete_space G'] - -- H for a normed group (hypotheses of mem_ℒp) - [normed_add_comm_group H] - -section Lp_meas - -/-! ## The subset `Lp_meas` of `Lp` functions a.e. measurable with respect to a sub-sigma-algebra -/ - -variables (F) - -/-- `Lp_meas_subgroup F m p μ` is the subspace of `Lp F p μ` containing functions `f` verifying -`ae_strongly_measurable' m f μ`, i.e. functions which are `μ`-a.e. equal to -an `m`-strongly measurable function. -/ -def Lp_meas_subgroup (m : measurable_space α) [measurable_space α] (p : ℝ≥0∞) (μ : measure α) : - add_subgroup (Lp F p μ) := -{ carrier := {f : (Lp F p μ) | ae_strongly_measurable' m f μ} , - zero_mem' := ⟨(0 : α → F), @strongly_measurable_zero _ _ m _ _, Lp.coe_fn_zero _ _ _⟩, - add_mem' := λ f g hf hg, (hf.add hg).congr (Lp.coe_fn_add f g).symm, - neg_mem' := λ f hf, ae_strongly_measurable'.congr hf.neg (Lp.coe_fn_neg f).symm, } - -variables (𝕜) -/-- `Lp_meas F 𝕜 m p μ` is the subspace of `Lp F p μ` containing functions `f` verifying -`ae_strongly_measurable' m f μ`, i.e. functions which are `μ`-a.e. equal to -an `m`-strongly measurable function. -/ -def Lp_meas (m : measurable_space α) [measurable_space α] (p : ℝ≥0∞) - (μ : measure α) : - submodule 𝕜 (Lp F p μ) := -{ carrier := {f : (Lp F p μ) | ae_strongly_measurable' m f μ} , - zero_mem' := ⟨(0 : α → F), @strongly_measurable_zero _ _ m _ _, Lp.coe_fn_zero _ _ _⟩, - add_mem' := λ f g hf hg, (hf.add hg).congr (Lp.coe_fn_add f g).symm, - smul_mem' := λ c f hf, (hf.const_smul c).congr (Lp.coe_fn_smul c f).symm, } -variables {F 𝕜} - -variables - -lemma mem_Lp_meas_subgroup_iff_ae_strongly_measurable' {m m0 : measurable_space α} {μ : measure α} - {f : Lp F p μ} : - f ∈ Lp_meas_subgroup F m p μ ↔ ae_strongly_measurable' m f μ := -by rw [← add_subgroup.mem_carrier, Lp_meas_subgroup, set.mem_set_of_eq] - -lemma mem_Lp_meas_iff_ae_strongly_measurable' - {m m0 : measurable_space α} {μ : measure α} {f : Lp F p μ} : - f ∈ Lp_meas F 𝕜 m p μ ↔ ae_strongly_measurable' m f μ := -by rw [← set_like.mem_coe, ← submodule.mem_carrier, Lp_meas, set.mem_set_of_eq] - -lemma Lp_meas.ae_strongly_measurable' - {m m0 : measurable_space α} {μ : measure α} (f : Lp_meas F 𝕜 m p μ) : - ae_strongly_measurable' m f μ := -mem_Lp_meas_iff_ae_strongly_measurable'.mp f.mem - -lemma mem_Lp_meas_self - {m0 : measurable_space α} (μ : measure α) (f : Lp F p μ) : - f ∈ Lp_meas F 𝕜 m0 p μ := -mem_Lp_meas_iff_ae_strongly_measurable'.mpr (Lp.ae_strongly_measurable f) - -lemma Lp_meas_subgroup_coe {m m0 : measurable_space α} {μ : measure α} - {f : Lp_meas_subgroup F m p μ} : - ⇑f = (f : Lp F p μ) := -coe_fn_coe_base f - -lemma Lp_meas_coe {m m0 : measurable_space α} {μ : measure α} {f : Lp_meas F 𝕜 m p μ} : - ⇑f = (f : Lp F p μ) := -coe_fn_coe_base f - -lemma mem_Lp_meas_indicator_const_Lp {m m0 : measurable_space α} (hm : m ≤ m0) - {μ : measure α} {s : set α} (hs : measurable_set[m] s) (hμs : μ s ≠ ∞) {c : F} : - indicator_const_Lp p (hm s hs) hμs c ∈ Lp_meas F 𝕜 m p μ := -⟨s.indicator (λ x : α, c), (@strongly_measurable_const _ _ m _ _).indicator hs, - indicator_const_Lp_coe_fn⟩ - -section complete_subspace - -/-! ## The subspace `Lp_meas` is complete. - -We define an `isometry_equiv` between `Lp_meas_subgroup` and the `Lp` space corresponding to the -measure `μ.trim hm`. As a consequence, the completeness of `Lp` implies completeness of -`Lp_meas_subgroup` (and `Lp_meas`). -/ - -variables {ι : Type*} {m m0 : measurable_space α} {μ : measure α} - -/-- If `f` belongs to `Lp_meas_subgroup F m p μ`, then the measurable function it is almost -everywhere equal to (given by `ae_measurable.mk`) belongs to `ℒp` for the measure `μ.trim hm`. -/ -lemma mem_ℒp_trim_of_mem_Lp_meas_subgroup (hm : m ≤ m0) (f : Lp F p μ) - (hf_meas : f ∈ Lp_meas_subgroup F m p μ) : - mem_ℒp (mem_Lp_meas_subgroup_iff_ae_strongly_measurable'.mp hf_meas).some p (μ.trim hm) := -begin - have hf : ae_strongly_measurable' m f μ, - from (mem_Lp_meas_subgroup_iff_ae_strongly_measurable'.mp hf_meas), - let g := hf.some, - obtain ⟨hg, hfg⟩ := hf.some_spec, - change mem_ℒp g p (μ.trim hm), - refine ⟨hg.ae_strongly_measurable, _⟩, - have h_snorm_fg : snorm g p (μ.trim hm) = snorm f p μ, - by { rw snorm_trim hm hg, exact snorm_congr_ae hfg.symm, }, - rw h_snorm_fg, - exact Lp.snorm_lt_top f, -end - -/-- If `f` belongs to `Lp` for the measure `μ.trim hm`, then it belongs to the subgroup -`Lp_meas_subgroup F m p μ`. -/ -lemma mem_Lp_meas_subgroup_to_Lp_of_trim (hm : m ≤ m0) (f : Lp F p (μ.trim hm)) : - (mem_ℒp_of_mem_ℒp_trim hm (Lp.mem_ℒp f)).to_Lp f ∈ Lp_meas_subgroup F m p μ := -begin - let hf_mem_ℒp := mem_ℒp_of_mem_ℒp_trim hm (Lp.mem_ℒp f), - rw mem_Lp_meas_subgroup_iff_ae_strongly_measurable', - refine ae_strongly_measurable'.congr _ (mem_ℒp.coe_fn_to_Lp hf_mem_ℒp).symm, - refine ae_strongly_measurable'_of_ae_strongly_measurable'_trim hm _, - exact Lp.ae_strongly_measurable f, -end - -variables (F p μ) -/-- Map from `Lp_meas_subgroup` to `Lp F p (μ.trim hm)`. -/ -def Lp_meas_subgroup_to_Lp_trim (hm : m ≤ m0) (f : Lp_meas_subgroup F m p μ) : Lp F p (μ.trim hm) := -mem_ℒp.to_Lp (mem_Lp_meas_subgroup_iff_ae_strongly_measurable'.mp f.mem).some - (mem_ℒp_trim_of_mem_Lp_meas_subgroup hm f f.mem) - -variables (𝕜) -/-- Map from `Lp_meas` to `Lp F p (μ.trim hm)`. -/ -def Lp_meas_to_Lp_trim (hm : m ≤ m0) (f : Lp_meas F 𝕜 m p μ) : Lp F p (μ.trim hm) := -mem_ℒp.to_Lp (mem_Lp_meas_iff_ae_strongly_measurable'.mp f.mem).some - (mem_ℒp_trim_of_mem_Lp_meas_subgroup hm f f.mem) -variables {𝕜} - -/-- Map from `Lp F p (μ.trim hm)` to `Lp_meas_subgroup`, inverse of -`Lp_meas_subgroup_to_Lp_trim`. -/ -def Lp_trim_to_Lp_meas_subgroup (hm : m ≤ m0) (f : Lp F p (μ.trim hm)) : Lp_meas_subgroup F m p μ := -⟨(mem_ℒp_of_mem_ℒp_trim hm (Lp.mem_ℒp f)).to_Lp f, mem_Lp_meas_subgroup_to_Lp_of_trim hm f⟩ - -variables (𝕜) -/-- Map from `Lp F p (μ.trim hm)` to `Lp_meas`, inverse of `Lp_meas_to_Lp_trim`. -/ -def Lp_trim_to_Lp_meas (hm : m ≤ m0) (f : Lp F p (μ.trim hm)) : Lp_meas F 𝕜 m p μ := -⟨(mem_ℒp_of_mem_ℒp_trim hm (Lp.mem_ℒp f)).to_Lp f, mem_Lp_meas_subgroup_to_Lp_of_trim hm f⟩ - -variables {F 𝕜 p μ} - -lemma Lp_meas_subgroup_to_Lp_trim_ae_eq (hm : m ≤ m0) (f : Lp_meas_subgroup F m p μ) : - Lp_meas_subgroup_to_Lp_trim F p μ hm f =ᵐ[μ] f := -(ae_eq_of_ae_eq_trim (mem_ℒp.coe_fn_to_Lp (mem_ℒp_trim_of_mem_Lp_meas_subgroup hm ↑f f.mem))).trans - (mem_Lp_meas_subgroup_iff_ae_strongly_measurable'.mp f.mem).some_spec.2.symm - -lemma Lp_trim_to_Lp_meas_subgroup_ae_eq (hm : m ≤ m0) (f : Lp F p (μ.trim hm)) : - Lp_trim_to_Lp_meas_subgroup F p μ hm f =ᵐ[μ] f := -mem_ℒp.coe_fn_to_Lp _ - -lemma Lp_meas_to_Lp_trim_ae_eq (hm : m ≤ m0) (f : Lp_meas F 𝕜 m p μ) : - Lp_meas_to_Lp_trim F 𝕜 p μ hm f =ᵐ[μ] f := -(ae_eq_of_ae_eq_trim (mem_ℒp.coe_fn_to_Lp (mem_ℒp_trim_of_mem_Lp_meas_subgroup hm ↑f f.mem))).trans - (mem_Lp_meas_subgroup_iff_ae_strongly_measurable'.mp f.mem).some_spec.2.symm - -lemma Lp_trim_to_Lp_meas_ae_eq (hm : m ≤ m0) (f : Lp F p (μ.trim hm)) : - Lp_trim_to_Lp_meas F 𝕜 p μ hm f =ᵐ[μ] f := -mem_ℒp.coe_fn_to_Lp _ - -/-- `Lp_trim_to_Lp_meas_subgroup` is a right inverse of `Lp_meas_subgroup_to_Lp_trim`. -/ -lemma Lp_meas_subgroup_to_Lp_trim_right_inv (hm : m ≤ m0) : - function.right_inverse (Lp_trim_to_Lp_meas_subgroup F p μ hm) - (Lp_meas_subgroup_to_Lp_trim F p μ hm) := -begin - intro f, - ext1, - refine ae_eq_trim_of_strongly_measurable hm - (Lp.strongly_measurable _) (Lp.strongly_measurable _) _, - exact (Lp_meas_subgroup_to_Lp_trim_ae_eq hm _).trans (Lp_trim_to_Lp_meas_subgroup_ae_eq hm _), -end - -/-- `Lp_trim_to_Lp_meas_subgroup` is a left inverse of `Lp_meas_subgroup_to_Lp_trim`. -/ -lemma Lp_meas_subgroup_to_Lp_trim_left_inv (hm : m ≤ m0) : - function.left_inverse (Lp_trim_to_Lp_meas_subgroup F p μ hm) - (Lp_meas_subgroup_to_Lp_trim F p μ hm) := -begin - intro f, - ext1, - ext1, - rw ← Lp_meas_subgroup_coe, - exact (Lp_trim_to_Lp_meas_subgroup_ae_eq hm _).trans (Lp_meas_subgroup_to_Lp_trim_ae_eq hm _), -end - -lemma Lp_meas_subgroup_to_Lp_trim_add (hm : m ≤ m0) (f g : Lp_meas_subgroup F m p μ) : - Lp_meas_subgroup_to_Lp_trim F p μ hm (f + g) - = Lp_meas_subgroup_to_Lp_trim F p μ hm f + Lp_meas_subgroup_to_Lp_trim F p μ hm g := -begin - ext1, - refine eventually_eq.trans _ (Lp.coe_fn_add _ _).symm, - refine ae_eq_trim_of_strongly_measurable hm (Lp.strongly_measurable _) _ _, - { exact (Lp.strongly_measurable _).add (Lp.strongly_measurable _), }, - refine (Lp_meas_subgroup_to_Lp_trim_ae_eq hm _).trans _, - refine eventually_eq.trans _ - (eventually_eq.add (Lp_meas_subgroup_to_Lp_trim_ae_eq hm f).symm - (Lp_meas_subgroup_to_Lp_trim_ae_eq hm g).symm), - refine (Lp.coe_fn_add _ _).trans _, - simp_rw Lp_meas_subgroup_coe, - exact eventually_of_forall (λ x, by refl), -end - -lemma Lp_meas_subgroup_to_Lp_trim_neg (hm : m ≤ m0) (f : Lp_meas_subgroup F m p μ) : - Lp_meas_subgroup_to_Lp_trim F p μ hm (-f) - = -Lp_meas_subgroup_to_Lp_trim F p μ hm f := -begin - ext1, - refine eventually_eq.trans _ (Lp.coe_fn_neg _).symm, - refine ae_eq_trim_of_strongly_measurable hm (Lp.strongly_measurable _) _ _, - { exact @strongly_measurable.neg _ _ _ m _ _ _ (Lp.strongly_measurable _), }, - refine (Lp_meas_subgroup_to_Lp_trim_ae_eq hm _).trans _, - refine eventually_eq.trans _ - (eventually_eq.neg (Lp_meas_subgroup_to_Lp_trim_ae_eq hm f).symm), - refine (Lp.coe_fn_neg _).trans _, - simp_rw Lp_meas_subgroup_coe, - exact eventually_of_forall (λ x, by refl), -end - -lemma Lp_meas_subgroup_to_Lp_trim_sub (hm : m ≤ m0) (f g : Lp_meas_subgroup F m p μ) : - Lp_meas_subgroup_to_Lp_trim F p μ hm (f - g) - = Lp_meas_subgroup_to_Lp_trim F p μ hm f - Lp_meas_subgroup_to_Lp_trim F p μ hm g := -by rw [sub_eq_add_neg, sub_eq_add_neg, Lp_meas_subgroup_to_Lp_trim_add, - Lp_meas_subgroup_to_Lp_trim_neg] - -lemma Lp_meas_to_Lp_trim_smul (hm : m ≤ m0) (c : 𝕜) (f : Lp_meas F 𝕜 m p μ) : - Lp_meas_to_Lp_trim F 𝕜 p μ hm (c • f) = c • Lp_meas_to_Lp_trim F 𝕜 p μ hm f := -begin - ext1, - refine eventually_eq.trans _ (Lp.coe_fn_smul _ _).symm, - refine ae_eq_trim_of_strongly_measurable hm (Lp.strongly_measurable _) _ _, - { exact (Lp.strongly_measurable _).const_smul c, }, - refine (Lp_meas_to_Lp_trim_ae_eq hm _).trans _, - refine (Lp.coe_fn_smul _ _).trans _, - refine (Lp_meas_to_Lp_trim_ae_eq hm f).mono (λ x hx, _), - rw [pi.smul_apply, pi.smul_apply, hx], - refl, -end - -/-- `Lp_meas_subgroup_to_Lp_trim` preserves the norm. -/ -lemma Lp_meas_subgroup_to_Lp_trim_norm_map [hp : fact (1 ≤ p)] (hm : m ≤ m0) - (f : Lp_meas_subgroup F m p μ) : - ‖Lp_meas_subgroup_to_Lp_trim F p μ hm f‖ = ‖f‖ := -begin - rw [Lp.norm_def, snorm_trim hm (Lp.strongly_measurable _), - snorm_congr_ae (Lp_meas_subgroup_to_Lp_trim_ae_eq hm _), Lp_meas_subgroup_coe, ← Lp.norm_def], - congr, -end - -lemma isometry_Lp_meas_subgroup_to_Lp_trim [hp : fact (1 ≤ p)] (hm : m ≤ m0) : - isometry (Lp_meas_subgroup_to_Lp_trim F p μ hm) := -isometry.of_dist_eq $ λ f g, by rw [dist_eq_norm, ← Lp_meas_subgroup_to_Lp_trim_sub, - Lp_meas_subgroup_to_Lp_trim_norm_map, dist_eq_norm] - -variables (F p μ) -/-- `Lp_meas_subgroup` and `Lp F p (μ.trim hm)` are isometric. -/ -def Lp_meas_subgroup_to_Lp_trim_iso [hp : fact (1 ≤ p)] (hm : m ≤ m0) : - Lp_meas_subgroup F m p μ ≃ᵢ Lp F p (μ.trim hm) := -{ to_fun := Lp_meas_subgroup_to_Lp_trim F p μ hm, - inv_fun := Lp_trim_to_Lp_meas_subgroup F p μ hm, - left_inv := Lp_meas_subgroup_to_Lp_trim_left_inv hm, - right_inv := Lp_meas_subgroup_to_Lp_trim_right_inv hm, - isometry_to_fun := isometry_Lp_meas_subgroup_to_Lp_trim hm, } - -variables (𝕜) -/-- `Lp_meas_subgroup` and `Lp_meas` are isometric. -/ -def Lp_meas_subgroup_to_Lp_meas_iso [hp : fact (1 ≤ p)] : - Lp_meas_subgroup F m p μ ≃ᵢ Lp_meas F 𝕜 m p μ := -isometry_equiv.refl (Lp_meas_subgroup F m p μ) - -/-- `Lp_meas` and `Lp F p (μ.trim hm)` are isometric, with a linear equivalence. -/ -def Lp_meas_to_Lp_trim_lie [hp : fact (1 ≤ p)] (hm : m ≤ m0) : - Lp_meas F 𝕜 m p μ ≃ₗᵢ[𝕜] Lp F p (μ.trim hm) := -{ to_fun := Lp_meas_to_Lp_trim F 𝕜 p μ hm, - inv_fun := Lp_trim_to_Lp_meas F 𝕜 p μ hm, - left_inv := Lp_meas_subgroup_to_Lp_trim_left_inv hm, - right_inv := Lp_meas_subgroup_to_Lp_trim_right_inv hm, - map_add' := Lp_meas_subgroup_to_Lp_trim_add hm, - map_smul' := Lp_meas_to_Lp_trim_smul hm, - norm_map' := Lp_meas_subgroup_to_Lp_trim_norm_map hm, } -variables {F 𝕜 p μ} - -instance [hm : fact (m ≤ m0)] [complete_space F] [hp : fact (1 ≤ p)] : - complete_space (Lp_meas_subgroup F m p μ) := -by { rw (Lp_meas_subgroup_to_Lp_trim_iso F p μ hm.elim).complete_space_iff, apply_instance, } - --- For now just no-lint this; lean4's tree-based logging will make this easier to debug. --- One possible change might be to generalize `𝕜` from `is_R_or_C` to `normed_field`, as this --- result may well hold there. -@[nolint fails_quickly] -instance [hm : fact (m ≤ m0)] [complete_space F] [hp : fact (1 ≤ p)] : - complete_space (Lp_meas F 𝕜 m p μ) := -by { rw (Lp_meas_subgroup_to_Lp_meas_iso F 𝕜 p μ).symm.complete_space_iff, apply_instance, } - -lemma is_complete_ae_strongly_measurable' [hp : fact (1 ≤ p)] [complete_space F] (hm : m ≤ m0) : - is_complete {f : Lp F p μ | ae_strongly_measurable' m f μ} := -begin - rw ← complete_space_coe_iff_is_complete, - haveI : fact (m ≤ m0) := ⟨hm⟩, - change complete_space (Lp_meas_subgroup F m p μ), - apply_instance, -end - -lemma is_closed_ae_strongly_measurable' [hp : fact (1 ≤ p)] [complete_space F] (hm : m ≤ m0) : - is_closed {f : Lp F p μ | ae_strongly_measurable' m f μ} := -is_complete.is_closed (is_complete_ae_strongly_measurable' hm) - -end complete_subspace - -section strongly_measurable - -variables {m m0 : measurable_space α} {μ : measure α} - -/-- We do not get `ae_fin_strongly_measurable f (μ.trim hm)`, since we don't have -`f =ᵐ[μ.trim hm] Lp_meas_to_Lp_trim F 𝕜 p μ hm f` but only the weaker -`f =ᵐ[μ] Lp_meas_to_Lp_trim F 𝕜 p μ hm f`. -/ -lemma Lp_meas.ae_fin_strongly_measurable' (hm : m ≤ m0) (f : Lp_meas F 𝕜 m p μ) (hp_ne_zero : p ≠ 0) - (hp_ne_top : p ≠ ∞) : - ∃ g, fin_strongly_measurable g (μ.trim hm) ∧ f =ᵐ[μ] g := -⟨Lp_meas_subgroup_to_Lp_trim F p μ hm f, Lp.fin_strongly_measurable _ hp_ne_zero hp_ne_top, - (Lp_meas_subgroup_to_Lp_trim_ae_eq hm f).symm⟩ - -/-- When applying the inverse of `Lp_meas_to_Lp_trim_lie` (which takes a function in the Lp space of -the sub-sigma algebra and returns its version in the larger Lp space) to an indicator of the -sub-sigma-algebra, we obtain an indicator in the Lp space of the larger sigma-algebra. -/ -lemma Lp_meas_to_Lp_trim_lie_symm_indicator [one_le_p : fact (1 ≤ p)] [normed_space ℝ F] - {hm : m ≤ m0} {s : set α} {μ : measure α} - (hs : measurable_set[m] s) (hμs : μ.trim hm s ≠ ∞) (c : F) : - ((Lp_meas_to_Lp_trim_lie F ℝ p μ hm).symm - (indicator_const_Lp p hs hμs c) : Lp F p μ) - = indicator_const_Lp p (hm s hs) ((le_trim hm).trans_lt hμs.lt_top).ne c := -begin - ext1, - rw ← Lp_meas_coe, - change Lp_trim_to_Lp_meas F ℝ p μ hm (indicator_const_Lp p hs hμs c) - =ᵐ[μ] (indicator_const_Lp p _ _ c : α → F), - refine (Lp_trim_to_Lp_meas_ae_eq hm _).trans _, - exact (ae_eq_of_ae_eq_trim indicator_const_Lp_coe_fn).trans indicator_const_Lp_coe_fn.symm, -end - -lemma Lp_meas_to_Lp_trim_lie_symm_to_Lp [one_le_p : fact (1 ≤ p)] [normed_space ℝ F] - (hm : m ≤ m0) (f : α → F) (hf : mem_ℒp f p (μ.trim hm)) : - ((Lp_meas_to_Lp_trim_lie F ℝ p μ hm).symm (hf.to_Lp f) : Lp F p μ) - = (mem_ℒp_of_mem_ℒp_trim hm hf).to_Lp f := -begin - ext1, - rw ← Lp_meas_coe, - refine (Lp_trim_to_Lp_meas_ae_eq hm _).trans _, - exact (ae_eq_of_ae_eq_trim (mem_ℒp.coe_fn_to_Lp hf)).trans (mem_ℒp.coe_fn_to_Lp _).symm, -end - -end strongly_measurable - -end Lp_meas - - -section induction - -variables {m m0 : measurable_space α} {μ : measure α} [fact (1 ≤ p)] [normed_space ℝ F] - -/-- Auxiliary lemma for `Lp.induction_strongly_measurable`. -/ -@[elab_as_eliminator] -lemma Lp.induction_strongly_measurable_aux (hm : m ≤ m0) (hp_ne_top : p ≠ ∞) (P : Lp F p μ → Prop) - (h_ind : ∀ (c : F) {s : set α} (hs : measurable_set[m] s) (hμs : μ s < ∞), - P (Lp.simple_func.indicator_const p (hm s hs) hμs.ne c)) - (h_add : ∀ ⦃f g⦄, ∀ hf : mem_ℒp f p μ, ∀ hg : mem_ℒp g p μ, - ∀ hfm : ae_strongly_measurable' m f μ, ∀ hgm : ae_strongly_measurable' m g μ, - disjoint (function.support f) (function.support g) → - P (hf.to_Lp f) → P (hg.to_Lp g) → P ((hf.to_Lp f) + (hg.to_Lp g))) - (h_closed : is_closed {f : Lp_meas F ℝ m p μ | P f}) : - ∀ f : Lp F p μ, ae_strongly_measurable' m f μ → P f := -begin - intros f hf, - let f' := (⟨f, hf⟩ : Lp_meas F ℝ m p μ), - let g := Lp_meas_to_Lp_trim_lie F ℝ p μ hm f', - have hfg : f' = (Lp_meas_to_Lp_trim_lie F ℝ p μ hm).symm g, - by simp only [linear_isometry_equiv.symm_apply_apply], - change P ↑f', - rw hfg, - refine @Lp.induction α F m _ p (μ.trim hm) _ hp_ne_top - (λ g, P ((Lp_meas_to_Lp_trim_lie F ℝ p μ hm).symm g)) _ _ _ g, - { intros b t ht hμt, - rw [Lp.simple_func.coe_indicator_const, - Lp_meas_to_Lp_trim_lie_symm_indicator ht hμt.ne b], - have hμt' : μ t < ∞, from (le_trim hm).trans_lt hμt, - specialize h_ind b ht hμt', - rwa Lp.simple_func.coe_indicator_const at h_ind, }, - { intros f g hf hg h_disj hfP hgP, - rw linear_isometry_equiv.map_add, - push_cast, - have h_eq : ∀ (f : α → F) (hf : mem_ℒp f p (μ.trim hm)), - ((Lp_meas_to_Lp_trim_lie F ℝ p μ hm).symm (mem_ℒp.to_Lp f hf) : Lp F p μ) - = (mem_ℒp_of_mem_ℒp_trim hm hf).to_Lp f, - from Lp_meas_to_Lp_trim_lie_symm_to_Lp hm, - rw h_eq f hf at hfP ⊢, - rw h_eq g hg at hgP ⊢, - exact h_add (mem_ℒp_of_mem_ℒp_trim hm hf) (mem_ℒp_of_mem_ℒp_trim hm hg) - (ae_strongly_measurable'_of_ae_strongly_measurable'_trim hm hf.ae_strongly_measurable) - (ae_strongly_measurable'_of_ae_strongly_measurable'_trim hm hg.ae_strongly_measurable) - h_disj hfP hgP, }, - { change is_closed ((Lp_meas_to_Lp_trim_lie F ℝ p μ hm).symm ⁻¹' {g : Lp_meas F ℝ m p μ | P ↑g}), - exact is_closed.preimage (linear_isometry_equiv.continuous _) h_closed, }, -end - -/-- To prove something for an `Lp` function a.e. strongly measurable with respect to a -sub-σ-algebra `m` in a normed space, it suffices to show that -* the property holds for (multiples of) characteristic functions which are measurable w.r.t. `m`; -* is closed under addition; -* the set of functions in `Lp` strongly measurable w.r.t. `m` for which the property holds is - closed. --/ -@[elab_as_eliminator] -lemma Lp.induction_strongly_measurable (hm : m ≤ m0) (hp_ne_top : p ≠ ∞) (P : Lp F p μ → Prop) - (h_ind : ∀ (c : F) {s : set α} (hs : measurable_set[m] s) (hμs : μ s < ∞), - P (Lp.simple_func.indicator_const p (hm s hs) hμs.ne c)) - (h_add : ∀ ⦃f g⦄, ∀ hf : mem_ℒp f p μ, ∀ hg : mem_ℒp g p μ, - ∀ hfm : strongly_measurable[m] f, ∀ hgm : strongly_measurable[m] g, - disjoint (function.support f) (function.support g) → - P (hf.to_Lp f) → P (hg.to_Lp g) → P ((hf.to_Lp f) + (hg.to_Lp g))) - (h_closed : is_closed {f : Lp_meas F ℝ m p μ | P f}) : - ∀ f : Lp F p μ, ae_strongly_measurable' m f μ → P f := -begin - intros f hf, - suffices h_add_ae : ∀ ⦃f g⦄, ∀ hf : mem_ℒp f p μ, ∀ hg : mem_ℒp g p μ, - ∀ hfm : ae_strongly_measurable' m f μ, ∀ hgm : ae_strongly_measurable' m g μ, - disjoint (function.support f) (function.support g) → - P (hf.to_Lp f) → P (hg.to_Lp g) → P ((hf.to_Lp f) + (hg.to_Lp g)), - from Lp.induction_strongly_measurable_aux hm hp_ne_top P h_ind h_add_ae h_closed f hf, - intros f g hf hg hfm hgm h_disj hPf hPg, - let s_f : set α := function.support (hfm.mk f), - have hs_f : measurable_set[m] s_f := hfm.strongly_measurable_mk.measurable_set_support, - have hs_f_eq : s_f =ᵐ[μ] function.support f := hfm.ae_eq_mk.symm.support, - let s_g : set α := function.support (hgm.mk g), - have hs_g : measurable_set[m] s_g := hgm.strongly_measurable_mk.measurable_set_support, - have hs_g_eq : s_g =ᵐ[μ] function.support g := hgm.ae_eq_mk.symm.support, - have h_inter_empty : ((s_f ∩ s_g) : set α) =ᵐ[μ] (∅ : set α), - { refine (hs_f_eq.inter hs_g_eq).trans _, - suffices : function.support f ∩ function.support g = ∅, by rw this, - exact set.disjoint_iff_inter_eq_empty.mp h_disj, }, - let f' := (s_f \ s_g).indicator (hfm.mk f), - have hff' : f =ᵐ[μ] f', - { have : s_f \ s_g =ᵐ[μ] s_f, - { rw [← set.diff_inter_self_eq_diff, set.inter_comm], - refine ((ae_eq_refl s_f).diff h_inter_empty).trans _, - rw set.diff_empty, }, - refine ((indicator_ae_eq_of_ae_eq_set this).trans _).symm, - rw set.indicator_support, - exact hfm.ae_eq_mk.symm, }, - have hf'_meas : strongly_measurable[m] f', - from hfm.strongly_measurable_mk.indicator (hs_f.diff hs_g), - have hf'_Lp : mem_ℒp f' p μ := hf.ae_eq hff', - let g' := (s_g \ s_f).indicator (hgm.mk g), - have hgg' : g =ᵐ[μ] g', - { have : s_g \ s_f =ᵐ[μ] s_g, - { rw [← set.diff_inter_self_eq_diff], - refine ((ae_eq_refl s_g).diff h_inter_empty).trans _, - rw set.diff_empty, }, - refine ((indicator_ae_eq_of_ae_eq_set this).trans _).symm, - rw set.indicator_support, - exact hgm.ae_eq_mk.symm, }, - have hg'_meas : strongly_measurable[m] g', - from hgm.strongly_measurable_mk.indicator (hs_g.diff hs_f), - have hg'_Lp : mem_ℒp g' p μ := hg.ae_eq hgg', - have h_disj : disjoint (function.support f') (function.support g'), - { have : disjoint (s_f \ s_g) (s_g \ s_f) := disjoint_sdiff_sdiff, - exact this.mono set.support_indicator_subset set.support_indicator_subset, }, - rw ← mem_ℒp.to_Lp_congr hf'_Lp hf hff'.symm at ⊢ hPf, - rw ← mem_ℒp.to_Lp_congr hg'_Lp hg hgg'.symm at ⊢ hPg, - exact h_add hf'_Lp hg'_Lp hf'_meas hg'_meas h_disj hPf hPg, -end - -/-- To prove something for an arbitrary `mem_ℒp` function a.e. strongly measurable with respect -to a sub-σ-algebra `m` in a normed space, it suffices to show that -* the property holds for (multiples of) characteristic functions which are measurable w.r.t. `m`; -* is closed under addition; -* the set of functions in the `Lᵖ` space strongly measurable w.r.t. `m` for which the property - holds is closed. -* the property is closed under the almost-everywhere equal relation. --/ -@[elab_as_eliminator] -lemma mem_ℒp.induction_strongly_measurable (hm : m ≤ m0) (hp_ne_top : p ≠ ∞) - (P : (α → F) → Prop) - (h_ind : ∀ (c : F) ⦃s⦄, measurable_set[m] s → μ s < ∞ → P (s.indicator (λ _, c))) - (h_add : ∀ ⦃f g : α → F⦄, disjoint (function.support f) (function.support g) - → mem_ℒp f p μ → mem_ℒp g p μ → strongly_measurable[m] f → strongly_measurable[m] g → - P f → P g → P (f + g)) - (h_closed : is_closed {f : Lp_meas F ℝ m p μ | P f} ) - (h_ae : ∀ ⦃f g⦄, f =ᵐ[μ] g → mem_ℒp f p μ → P f → P g) : - ∀ ⦃f : α → F⦄ (hf : mem_ℒp f p μ) (hfm : ae_strongly_measurable' m f μ), P f := -begin - intros f hf hfm, - let f_Lp := hf.to_Lp f, - have hfm_Lp : ae_strongly_measurable' m f_Lp μ, from hfm.congr hf.coe_fn_to_Lp.symm, - refine h_ae (hf.coe_fn_to_Lp) (Lp.mem_ℒp _) _, - change P f_Lp, - refine Lp.induction_strongly_measurable hm hp_ne_top (λ f, P ⇑f) _ _ h_closed f_Lp hfm_Lp, - { intros c s hs hμs, - rw Lp.simple_func.coe_indicator_const, - refine h_ae (indicator_const_Lp_coe_fn).symm _ (h_ind c hs hμs), - exact mem_ℒp_indicator_const p (hm s hs) c (or.inr hμs.ne), }, - { intros f g hf_mem hg_mem hfm hgm h_disj hfP hgP, - have hfP' : P f := h_ae (hf_mem.coe_fn_to_Lp) (Lp.mem_ℒp _) hfP, - have hgP' : P g := h_ae (hg_mem.coe_fn_to_Lp) (Lp.mem_ℒp _) hgP, - specialize h_add h_disj hf_mem hg_mem hfm hgm hfP' hgP', - refine h_ae _ (hf_mem.add hg_mem) h_add, - exact ((hf_mem.coe_fn_to_Lp).symm.add (hg_mem.coe_fn_to_Lp).symm).trans - (Lp.coe_fn_add _ _).symm, }, -end - -end induction - - -section uniqueness_of_conditional_expectation - -/-! ## Uniqueness of the conditional expectation -/ - -variables {m m0 : measurable_space α} {μ : measure α} - -lemma Lp_meas.ae_eq_zero_of_forall_set_integral_eq_zero - (hm : m ≤ m0) (f : Lp_meas E' 𝕜 m p μ) (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) - (hf_int_finite : ∀ s, measurable_set[m] s → μ s < ∞ → integrable_on f s μ) - (hf_zero : ∀ s : set α, measurable_set[m] s → μ s < ∞ → ∫ x in s, f x ∂μ = 0) : - f =ᵐ[μ] 0 := -begin - obtain ⟨g, hg_sm, hfg⟩ := Lp_meas.ae_fin_strongly_measurable' hm f hp_ne_zero hp_ne_top, - refine hfg.trans _, - refine ae_eq_zero_of_forall_set_integral_eq_of_fin_strongly_measurable_trim hm _ _ hg_sm, - { intros s hs hμs, - have hfg_restrict : f =ᵐ[μ.restrict s] g, from ae_restrict_of_ae hfg, - rw [integrable_on, integrable_congr hfg_restrict.symm], - exact hf_int_finite s hs hμs, }, - { intros s hs hμs, - have hfg_restrict : f =ᵐ[μ.restrict s] g, from ae_restrict_of_ae hfg, - rw integral_congr_ae hfg_restrict.symm, - exact hf_zero s hs hμs, }, -end - -include 𝕜 -variables (𝕜) - -lemma Lp.ae_eq_zero_of_forall_set_integral_eq_zero' - (hm : m ≤ m0) (f : Lp E' p μ) (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) - (hf_int_finite : ∀ s, measurable_set[m] s → μ s < ∞ → integrable_on f s μ) - (hf_zero : ∀ s : set α, measurable_set[m] s → μ s < ∞ → ∫ x in s, f x ∂μ = 0) - (hf_meas : ae_strongly_measurable' m f μ) : - f =ᵐ[μ] 0 := -begin - let f_meas : Lp_meas E' 𝕜 m p μ := ⟨f, hf_meas⟩, - have hf_f_meas : f =ᵐ[μ] f_meas, by simp only [coe_fn_coe_base', subtype.coe_mk], - refine hf_f_meas.trans _, - refine Lp_meas.ae_eq_zero_of_forall_set_integral_eq_zero hm f_meas hp_ne_zero hp_ne_top _ _, - { intros s hs hμs, - have hfg_restrict : f =ᵐ[μ.restrict s] f_meas, from ae_restrict_of_ae hf_f_meas, - rw [integrable_on, integrable_congr hfg_restrict.symm], - exact hf_int_finite s hs hμs, }, - { intros s hs hμs, - have hfg_restrict : f =ᵐ[μ.restrict s] f_meas, from ae_restrict_of_ae hf_f_meas, - rw integral_congr_ae hfg_restrict.symm, - exact hf_zero s hs hμs, }, -end - -/-- **Uniqueness of the conditional expectation** -/ -lemma Lp.ae_eq_of_forall_set_integral_eq' - (hm : m ≤ m0) (f g : Lp E' p μ) (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) - (hf_int_finite : ∀ s, measurable_set[m] s → μ s < ∞ → integrable_on f s μ) - (hg_int_finite : ∀ s, measurable_set[m] s → μ s < ∞ → integrable_on g s μ) - (hfg : ∀ s : set α, measurable_set[m] s → μ s < ∞ → ∫ x in s, f x ∂μ = ∫ x in s, g x ∂μ) - (hf_meas : ae_strongly_measurable' m f μ) (hg_meas : ae_strongly_measurable' m g μ) : - f =ᵐ[μ] g := -begin - suffices h_sub : ⇑(f-g) =ᵐ[μ] 0, - by { rw ← sub_ae_eq_zero, exact (Lp.coe_fn_sub f g).symm.trans h_sub, }, - have hfg' : ∀ s : set α, measurable_set[m] s → μ s < ∞ → ∫ x in s, (f - g) x ∂μ = 0, - { intros s hs hμs, - rw integral_congr_ae (ae_restrict_of_ae (Lp.coe_fn_sub f g)), - rw integral_sub' (hf_int_finite s hs hμs) (hg_int_finite s hs hμs), - exact sub_eq_zero.mpr (hfg s hs hμs), }, - have hfg_int : ∀ s, measurable_set[m] s → μ s < ∞ → integrable_on ⇑(f-g) s μ, - { intros s hs hμs, - rw [integrable_on, integrable_congr (ae_restrict_of_ae (Lp.coe_fn_sub f g))], - exact (hf_int_finite s hs hμs).sub (hg_int_finite s hs hμs), }, - have hfg_meas : ae_strongly_measurable' m ⇑(f - g) μ, - from ae_strongly_measurable'.congr (hf_meas.sub hg_meas) (Lp.coe_fn_sub f g).symm, - exact Lp.ae_eq_zero_of_forall_set_integral_eq_zero' 𝕜 hm (f-g) hp_ne_zero hp_ne_top hfg_int hfg' - hfg_meas, -end - -variables {𝕜} -omit 𝕜 - -lemma ae_eq_of_forall_set_integral_eq_of_sigma_finite' (hm : m ≤ m0) [sigma_finite (μ.trim hm)] - {f g : α → F'} - (hf_int_finite : ∀ s, measurable_set[m] s → μ s < ∞ → integrable_on f s μ) - (hg_int_finite : ∀ s, measurable_set[m] s → μ s < ∞ → integrable_on g s μ) - (hfg_eq : ∀ s : set α, measurable_set[m] s → μ s < ∞ → ∫ x in s, f x ∂μ = ∫ x in s, g x ∂μ) - (hfm : ae_strongly_measurable' m f μ) (hgm : ae_strongly_measurable' m g μ) : - f =ᵐ[μ] g := -begin - rw ← ae_eq_trim_iff_of_ae_strongly_measurable' hm hfm hgm, - have hf_mk_int_finite : ∀ s, measurable_set[m] s → μ.trim hm s < ∞ → - @integrable_on _ _ m _ (hfm.mk f) s (μ.trim hm), - { intros s hs hμs, - rw trim_measurable_set_eq hm hs at hμs, - rw [integrable_on, restrict_trim hm _ hs], - refine integrable.trim hm _ hfm.strongly_measurable_mk, - exact integrable.congr (hf_int_finite s hs hμs) (ae_restrict_of_ae hfm.ae_eq_mk), }, - have hg_mk_int_finite : ∀ s, measurable_set[m] s → μ.trim hm s < ∞ → - @integrable_on _ _ m _ (hgm.mk g) s (μ.trim hm), - { intros s hs hμs, - rw trim_measurable_set_eq hm hs at hμs, - rw [integrable_on, restrict_trim hm _ hs], - refine integrable.trim hm _ hgm.strongly_measurable_mk, - exact integrable.congr (hg_int_finite s hs hμs) (ae_restrict_of_ae hgm.ae_eq_mk), }, - have hfg_mk_eq : ∀ s : set α, measurable_set[m] s → μ.trim hm s < ∞ → - ∫ x in s, (hfm.mk f x) ∂(μ.trim hm) = ∫ x in s, (hgm.mk g x) ∂(μ.trim hm), - { intros s hs hμs, - rw trim_measurable_set_eq hm hs at hμs, - rw [restrict_trim hm _ hs, ← integral_trim hm hfm.strongly_measurable_mk, - ← integral_trim hm hgm.strongly_measurable_mk, - integral_congr_ae (ae_restrict_of_ae hfm.ae_eq_mk.symm), - integral_congr_ae (ae_restrict_of_ae hgm.ae_eq_mk.symm)], - exact hfg_eq s hs hμs, }, - exact ae_eq_of_forall_set_integral_eq_of_sigma_finite hf_mk_int_finite hg_mk_int_finite hfg_mk_eq, -end - -end uniqueness_of_conditional_expectation - - -section integral_norm_le - -variables {m m0 : measurable_space α} {μ : measure α} {s : set α} - -/-- Let `m` be a sub-σ-algebra of `m0`, `f` a `m0`-measurable function and `g` a `m`-measurable -function, such that their integrals coincide on `m`-measurable sets with finite measure. -Then `∫ x in s, ‖g x‖ ∂μ ≤ ∫ x in s, ‖f x‖ ∂μ` on all `m`-measurable sets with finite measure. -/ -lemma integral_norm_le_of_forall_fin_meas_integral_eq (hm : m ≤ m0) {f g : α → ℝ} - (hf : strongly_measurable f) (hfi : integrable_on f s μ) - (hg : strongly_measurable[m] g) (hgi : integrable_on g s μ) - (hgf : ∀ t, measurable_set[m] t → μ t < ∞ → ∫ x in t, g x ∂μ = ∫ x in t, f x ∂μ) - (hs : measurable_set[m] s) (hμs : μ s ≠ ∞) : - ∫ x in s, ‖g x‖ ∂μ ≤ ∫ x in s, ‖f x‖ ∂μ := -begin - rw [integral_norm_eq_pos_sub_neg hgi, integral_norm_eq_pos_sub_neg hfi], - have h_meas_nonneg_g : measurable_set[m] {x | 0 ≤ g x}, - from (@strongly_measurable_const _ _ m _ _).measurable_set_le hg, - have h_meas_nonneg_f : measurable_set {x | 0 ≤ f x}, - from strongly_measurable_const.measurable_set_le hf, - have h_meas_nonpos_g : measurable_set[m] {x | g x ≤ 0}, - from hg.measurable_set_le (@strongly_measurable_const _ _ m _ _), - have h_meas_nonpos_f : measurable_set {x | f x ≤ 0}, - from hf.measurable_set_le strongly_measurable_const, - refine sub_le_sub _ _, - { rw [measure.restrict_restrict (hm _ h_meas_nonneg_g), - measure.restrict_restrict h_meas_nonneg_f, - hgf _ (@measurable_set.inter α m _ _ h_meas_nonneg_g hs) - ((measure_mono (set.inter_subset_right _ _)).trans_lt (lt_top_iff_ne_top.mpr hμs)), - ← measure.restrict_restrict (hm _ h_meas_nonneg_g), - ← measure.restrict_restrict h_meas_nonneg_f], - exact set_integral_le_nonneg (hm _ h_meas_nonneg_g) hf hfi, }, - { rw [measure.restrict_restrict (hm _ h_meas_nonpos_g), - measure.restrict_restrict h_meas_nonpos_f, - hgf _ (@measurable_set.inter α m _ _ h_meas_nonpos_g hs) - ((measure_mono (set.inter_subset_right _ _)).trans_lt (lt_top_iff_ne_top.mpr hμs)), - ← measure.restrict_restrict (hm _ h_meas_nonpos_g), - ← measure.restrict_restrict h_meas_nonpos_f], - exact set_integral_nonpos_le (hm _ h_meas_nonpos_g) hf hfi, }, -end - -/-- Let `m` be a sub-σ-algebra of `m0`, `f` a `m0`-measurable function and `g` a `m`-measurable -function, such that their integrals coincide on `m`-measurable sets with finite measure. -Then `∫⁻ x in s, ‖g x‖₊ ∂μ ≤ ∫⁻ x in s, ‖f x‖₊ ∂μ` on all `m`-measurable sets with finite -measure. -/ -lemma lintegral_nnnorm_le_of_forall_fin_meas_integral_eq (hm : m ≤ m0) {f g : α → ℝ} - (hf : strongly_measurable f) (hfi : integrable_on f s μ) - (hg : strongly_measurable[m] g) (hgi : integrable_on g s μ) - (hgf : ∀ t, measurable_set[m] t → μ t < ∞ → ∫ x in t, g x ∂μ = ∫ x in t, f x ∂μ) - (hs : measurable_set[m] s) (hμs : μ s ≠ ∞) : - ∫⁻ x in s, ‖g x‖₊ ∂μ ≤ ∫⁻ x in s, ‖f x‖₊ ∂μ := -begin - rw [← of_real_integral_norm_eq_lintegral_nnnorm hfi, - ← of_real_integral_norm_eq_lintegral_nnnorm hgi, ennreal.of_real_le_of_real_iff], - { exact integral_norm_le_of_forall_fin_meas_integral_eq hm hf hfi hg hgi hgf hs hμs, }, - { exact integral_nonneg (λ x, norm_nonneg _), }, -end - -end integral_norm_le - -/-! ## Conditional expectation in L2 - -We define a conditional expectation in `L2`: it is the orthogonal projection on the subspace -`Lp_meas`. -/ - -section condexp_L2 - -variables [complete_space E] {m m0 : measurable_space α} {μ : measure α} - {s t : set α} - -local notation `⟪`x`, `y`⟫` := @inner 𝕜 E _ x y -local notation `⟪`x`, `y`⟫₂` := @inner 𝕜 (α →₂[μ] E) _ x y - -variables (𝕜) -/-- Conditional expectation of a function in L2 with respect to a sigma-algebra -/ -def condexp_L2 (hm : m ≤ m0) : (α →₂[μ] E) →L[𝕜] (Lp_meas E 𝕜 m 2 μ) := -@orthogonal_projection 𝕜 (α →₂[μ] E) _ _ _ (Lp_meas E 𝕜 m 2 μ) - (by { haveI : fact (m ≤ m0) := ⟨hm⟩, exact infer_instance, }) -variables {𝕜} - -lemma ae_strongly_measurable'_condexp_L2 (hm : m ≤ m0) (f : α →₂[μ] E) : - ae_strongly_measurable' m (condexp_L2 𝕜 hm f) μ := -Lp_meas.ae_strongly_measurable' _ - -lemma integrable_on_condexp_L2_of_measure_ne_top (hm : m ≤ m0) (hμs : μ s ≠ ∞) (f : α →₂[μ] E) : - integrable_on (condexp_L2 𝕜 hm f) s μ := -integrable_on_Lp_of_measure_ne_top ((condexp_L2 𝕜 hm f) : α →₂[μ] E) - fact_one_le_two_ennreal.elim hμs - -lemma integrable_condexp_L2_of_is_finite_measure (hm : m ≤ m0) [is_finite_measure μ] - {f : α →₂[μ] E} : - integrable (condexp_L2 𝕜 hm f) μ := -integrable_on_univ.mp $ integrable_on_condexp_L2_of_measure_ne_top hm (measure_ne_top _ _) f - -lemma norm_condexp_L2_le_one (hm : m ≤ m0) : ‖@condexp_L2 α E 𝕜 _ _ _ _ _ _ μ hm‖ ≤ 1 := -by { haveI : fact (m ≤ m0) := ⟨hm⟩, exact orthogonal_projection_norm_le _, } - -lemma norm_condexp_L2_le (hm : m ≤ m0) (f : α →₂[μ] E) : ‖condexp_L2 𝕜 hm f‖ ≤ ‖f‖ := -((@condexp_L2 _ E 𝕜 _ _ _ _ _ _ μ hm).le_op_norm f).trans - (mul_le_of_le_one_left (norm_nonneg _) (norm_condexp_L2_le_one hm)) - -lemma snorm_condexp_L2_le (hm : m ≤ m0) (f : α →₂[μ] E) : - snorm (condexp_L2 𝕜 hm f) 2 μ ≤ snorm f 2 μ := -begin - rw [Lp_meas_coe, ← ennreal.to_real_le_to_real (Lp.snorm_ne_top _) (Lp.snorm_ne_top _), - ← Lp.norm_def, ← Lp.norm_def, submodule.norm_coe], - exact norm_condexp_L2_le hm f, -end - -lemma norm_condexp_L2_coe_le (hm : m ≤ m0) (f : α →₂[μ] E) : - ‖(condexp_L2 𝕜 hm f : α →₂[μ] E)‖ ≤ ‖f‖ := -begin - rw [Lp.norm_def, Lp.norm_def, ← Lp_meas_coe], - refine (ennreal.to_real_le_to_real _ (Lp.snorm_ne_top _)).mpr (snorm_condexp_L2_le hm f), - exact Lp.snorm_ne_top _, -end - -lemma inner_condexp_L2_left_eq_right (hm : m ≤ m0) {f g : α →₂[μ] E} : - ⟪(condexp_L2 𝕜 hm f : α →₂[μ] E), g⟫₂ = ⟪f, (condexp_L2 𝕜 hm g : α →₂[μ] E)⟫₂ := -by { haveI : fact (m ≤ m0) := ⟨hm⟩, exact inner_orthogonal_projection_left_eq_right _ f g, } - -lemma condexp_L2_indicator_of_measurable (hm : m ≤ m0) - (hs : measurable_set[m] s) (hμs : μ s ≠ ∞) (c : E) : - (condexp_L2 𝕜 hm (indicator_const_Lp 2 (hm s hs) hμs c) : α →₂[μ] E) - = indicator_const_Lp 2 (hm s hs) hμs c := -begin - rw condexp_L2, - haveI : fact (m ≤ m0) := ⟨hm⟩, - have h_mem : indicator_const_Lp 2 (hm s hs) hμs c ∈ Lp_meas E 𝕜 m 2 μ, - from mem_Lp_meas_indicator_const_Lp hm hs hμs, - let ind := (⟨indicator_const_Lp 2 (hm s hs) hμs c, h_mem⟩ : Lp_meas E 𝕜 m 2 μ), - have h_coe_ind : (ind : α →₂[μ] E) = indicator_const_Lp 2 (hm s hs) hμs c, by refl, - have h_orth_mem := orthogonal_projection_mem_subspace_eq_self ind, - rw [← h_coe_ind, h_orth_mem], -end - -lemma inner_condexp_L2_eq_inner_fun (hm : m ≤ m0) (f g : α →₂[μ] E) - (hg : ae_strongly_measurable' m g μ) : - ⟪(condexp_L2 𝕜 hm f : α →₂[μ] E), g⟫₂ = ⟪f, g⟫₂ := -begin - symmetry, - rw [← sub_eq_zero, ← inner_sub_left, condexp_L2], - simp only [mem_Lp_meas_iff_ae_strongly_measurable'.mpr hg, orthogonal_projection_inner_eq_zero], -end - -section real - -variables {hm : m ≤ m0} - -lemma integral_condexp_L2_eq_of_fin_meas_real (f : Lp 𝕜 2 μ) (hs : measurable_set[m] s) - (hμs : μ s ≠ ∞) : - ∫ x in s, condexp_L2 𝕜 hm f x ∂μ = ∫ x in s, f x ∂μ := -begin - rw ← L2.inner_indicator_const_Lp_one (hm s hs) hμs, - have h_eq_inner : ∫ x in s, condexp_L2 𝕜 hm f x ∂μ - = inner (indicator_const_Lp 2 (hm s hs) hμs (1 : 𝕜)) (condexp_L2 𝕜 hm f), - { rw L2.inner_indicator_const_Lp_one (hm s hs) hμs, - congr, }, - rw [h_eq_inner, ← inner_condexp_L2_left_eq_right, condexp_L2_indicator_of_measurable hm hs hμs], -end - -lemma lintegral_nnnorm_condexp_L2_le (hs : measurable_set[m] s) (hμs : μ s ≠ ∞) (f : Lp ℝ 2 μ) : - ∫⁻ x in s, ‖condexp_L2 ℝ hm f x‖₊ ∂μ ≤ ∫⁻ x in s, ‖f x‖₊ ∂μ := -begin - let h_meas := Lp_meas.ae_strongly_measurable' (condexp_L2 ℝ hm f), - let g := h_meas.some, - have hg_meas : strongly_measurable[m] g, from h_meas.some_spec.1, - have hg_eq : g =ᵐ[μ] condexp_L2 ℝ hm f, from h_meas.some_spec.2.symm, - have hg_eq_restrict : g =ᵐ[μ.restrict s] condexp_L2 ℝ hm f, from ae_restrict_of_ae hg_eq, - have hg_nnnorm_eq : (λ x, (‖g x‖₊ : ℝ≥0∞)) - =ᵐ[μ.restrict s] (λ x, (‖condexp_L2 ℝ hm f x‖₊ : ℝ≥0∞)), - { refine hg_eq_restrict.mono (λ x hx, _), - dsimp only, - rw hx, }, - rw lintegral_congr_ae hg_nnnorm_eq.symm, - refine lintegral_nnnorm_le_of_forall_fin_meas_integral_eq hm - (Lp.strongly_measurable f) _ _ _ _ hs hμs, - { exact integrable_on_Lp_of_measure_ne_top f fact_one_le_two_ennreal.elim hμs, }, - { exact hg_meas, }, - { rw [integrable_on, integrable_congr hg_eq_restrict], - exact integrable_on_condexp_L2_of_measure_ne_top hm hμs f, }, - { intros t ht hμt, - rw ← integral_condexp_L2_eq_of_fin_meas_real f ht hμt.ne, - exact set_integral_congr_ae (hm t ht) (hg_eq.mono (λ x hx _, hx)), }, -end - -lemma condexp_L2_ae_eq_zero_of_ae_eq_zero (hs : measurable_set[m] s) (hμs : μ s ≠ ∞) - {f : Lp ℝ 2 μ} (hf : f =ᵐ[μ.restrict s] 0) : - condexp_L2 ℝ hm f =ᵐ[μ.restrict s] 0 := -begin - suffices h_nnnorm_eq_zero : ∫⁻ x in s, ‖condexp_L2 ℝ hm f x‖₊ ∂μ = 0, - { rw lintegral_eq_zero_iff at h_nnnorm_eq_zero, - refine h_nnnorm_eq_zero.mono (λ x hx, _), - dsimp only at hx, - rw pi.zero_apply at hx ⊢, - { rwa [ennreal.coe_eq_zero, nnnorm_eq_zero] at hx, }, - { refine measurable.coe_nnreal_ennreal (measurable.nnnorm _), - rw Lp_meas_coe, - exact (Lp.strongly_measurable _).measurable }, }, - refine le_antisymm _ (zero_le _), - refine (lintegral_nnnorm_condexp_L2_le hs hμs f).trans (le_of_eq _), - rw lintegral_eq_zero_iff, - { refine hf.mono (λ x hx, _), - dsimp only, - rw hx, - simp, }, - { exact (Lp.strongly_measurable _).ennnorm, }, -end - -lemma lintegral_nnnorm_condexp_L2_indicator_le_real - (hs : measurable_set s) (hμs : μ s ≠ ∞) (ht : measurable_set[m] t) (hμt : μ t ≠ ∞) : - ∫⁻ a in t, ‖condexp_L2 ℝ hm (indicator_const_Lp 2 hs hμs (1 : ℝ)) a‖₊ ∂μ ≤ μ (s ∩ t) := -begin - refine (lintegral_nnnorm_condexp_L2_le ht hμt _).trans (le_of_eq _), - have h_eq : ∫⁻ x in t, ‖(indicator_const_Lp 2 hs hμs (1 : ℝ)) x‖₊ ∂μ - = ∫⁻ x in t, s.indicator (λ x, (1 : ℝ≥0∞)) x ∂μ, - { 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], - simp only [one_mul, set.univ_inter, measurable_set.univ, measure.restrict_apply], -end - -end real - -/-- `condexp_L2` commutes with taking inner products with constants. See the lemma -`condexp_L2_comp_continuous_linear_map` for a more general result about commuting with continuous -linear maps. -/ -lemma condexp_L2_const_inner (hm : m ≤ m0) (f : Lp E 2 μ) (c : E) : - condexp_L2 𝕜 hm (((Lp.mem_ℒp f).const_inner c).to_Lp (λ a, ⟪c, f a⟫)) - =ᵐ[μ] λ a, ⟪c, condexp_L2 𝕜 hm f a⟫ := -begin - rw Lp_meas_coe, - have h_mem_Lp : mem_ℒp (λ a, ⟪c, condexp_L2 𝕜 hm f a⟫) 2 μ, - { refine mem_ℒp.const_inner _ _, rw Lp_meas_coe, exact Lp.mem_ℒp _, }, - have h_eq : h_mem_Lp.to_Lp _ =ᵐ[μ] λ a, ⟪c, condexp_L2 𝕜 hm f a⟫, from h_mem_Lp.coe_fn_to_Lp, - refine eventually_eq.trans _ h_eq, - refine Lp.ae_eq_of_forall_set_integral_eq' 𝕜 hm _ _ two_ne_zero ennreal.coe_ne_top - (λ s hs hμs, integrable_on_condexp_L2_of_measure_ne_top hm hμs.ne _) _ _ _ _, - { intros s hs hμs, - rw [integrable_on, integrable_congr (ae_restrict_of_ae h_eq)], - exact (integrable_on_condexp_L2_of_measure_ne_top hm hμs.ne _).const_inner _, }, - { intros s hs hμs, - rw [← Lp_meas_coe, integral_condexp_L2_eq_of_fin_meas_real _ hs hμs.ne, - integral_congr_ae (ae_restrict_of_ae h_eq), Lp_meas_coe, - ← L2.inner_indicator_const_Lp_eq_set_integral_inner 𝕜 ↑(condexp_L2 𝕜 hm f) (hm s hs) c hμs.ne, - ← inner_condexp_L2_left_eq_right, condexp_L2_indicator_of_measurable, - L2.inner_indicator_const_Lp_eq_set_integral_inner 𝕜 f (hm s hs) c hμs.ne, - set_integral_congr_ae (hm s hs) - ((mem_ℒp.coe_fn_to_Lp ((Lp.mem_ℒp f).const_inner c)).mono (λ x hx hxs, hx))], }, - { rw ← Lp_meas_coe, exact Lp_meas.ae_strongly_measurable' _, }, - { refine ae_strongly_measurable'.congr _ h_eq.symm, - exact (Lp_meas.ae_strongly_measurable' _).const_inner _, }, -end - -/-- `condexp_L2` verifies the equality of integrals defining the conditional expectation. -/ -lemma integral_condexp_L2_eq (hm : m ≤ m0) - (f : Lp E' 2 μ) (hs : measurable_set[m] s) (hμs : μ s ≠ ∞) : - ∫ x in s, condexp_L2 𝕜 hm f x ∂μ = ∫ x in s, f x ∂μ := -begin - rw [← sub_eq_zero, Lp_meas_coe, ← integral_sub' - (integrable_on_Lp_of_measure_ne_top _ fact_one_le_two_ennreal.elim hμs) - (integrable_on_Lp_of_measure_ne_top _ fact_one_le_two_ennreal.elim hμs)], - refine integral_eq_zero_of_forall_integral_inner_eq_zero 𝕜 _ _ _, - { rw integrable_congr (ae_restrict_of_ae (Lp.coe_fn_sub ↑(condexp_L2 𝕜 hm f) f).symm), - exact integrable_on_Lp_of_measure_ne_top _ fact_one_le_two_ennreal.elim hμs, }, - intro c, - simp_rw [pi.sub_apply, inner_sub_right], - rw integral_sub - ((integrable_on_Lp_of_measure_ne_top _ fact_one_le_two_ennreal.elim hμs).const_inner c) - ((integrable_on_Lp_of_measure_ne_top _ fact_one_le_two_ennreal.elim hμs).const_inner c), - have h_ae_eq_f := mem_ℒp.coe_fn_to_Lp ((Lp.mem_ℒp f).const_inner c), - rw [← Lp_meas_coe, sub_eq_zero, - ← set_integral_congr_ae (hm s hs) ((condexp_L2_const_inner hm f c).mono (λ x hx _, hx)), - ← set_integral_congr_ae (hm s hs) (h_ae_eq_f.mono (λ x hx _, hx))], - exact integral_condexp_L2_eq_of_fin_meas_real _ hs hμs, -end - -variables {E'' 𝕜' : Type*} [is_R_or_C 𝕜'] [normed_add_comm_group E''] - [inner_product_space 𝕜' E''] [complete_space E''] [normed_space ℝ E''] - -variables (𝕜 𝕜') -lemma condexp_L2_comp_continuous_linear_map (hm : m ≤ m0) (T : E' →L[ℝ] E'') (f : α →₂[μ] E') : - (condexp_L2 𝕜' hm (T.comp_Lp f) : α →₂[μ] E'') =ᵐ[μ] T.comp_Lp (condexp_L2 𝕜 hm f : α →₂[μ] E') := -begin - refine Lp.ae_eq_of_forall_set_integral_eq' 𝕜' hm _ _ two_ne_zero ennreal.coe_ne_top - (λ s hs hμs, integrable_on_condexp_L2_of_measure_ne_top hm hμs.ne _) - (λ s hs hμs, integrable_on_Lp_of_measure_ne_top _ fact_one_le_two_ennreal.elim hμs.ne) - _ _ _, - { intros s hs hμs, - rw [T.set_integral_comp_Lp _ (hm s hs), - T.integral_comp_comm - (integrable_on_Lp_of_measure_ne_top _ fact_one_le_two_ennreal.elim hμs.ne), - ← Lp_meas_coe, ← Lp_meas_coe, integral_condexp_L2_eq hm f hs hμs.ne, - integral_condexp_L2_eq hm (T.comp_Lp f) hs hμs.ne, T.set_integral_comp_Lp _ (hm s hs), - T.integral_comp_comm - (integrable_on_Lp_of_measure_ne_top f fact_one_le_two_ennreal.elim hμs.ne)], }, - { rw ← Lp_meas_coe, exact Lp_meas.ae_strongly_measurable' _, }, - { have h_coe := T.coe_fn_comp_Lp (condexp_L2 𝕜 hm f : α →₂[μ] E'), - rw ← eventually_eq at h_coe, - refine ae_strongly_measurable'.congr _ h_coe.symm, - exact (Lp_meas.ae_strongly_measurable' (condexp_L2 𝕜 hm f)).continuous_comp T.continuous, }, -end -variables {𝕜 𝕜'} - -section condexp_L2_indicator - -variables (𝕜) -lemma condexp_L2_indicator_ae_eq_smul (hm : m ≤ m0) (hs : measurable_set s) (hμs : μ s ≠ ∞) - (x : E') : - condexp_L2 𝕜 hm (indicator_const_Lp 2 hs hμs x) - =ᵐ[μ] λ a, (condexp_L2 ℝ hm (indicator_const_Lp 2 hs hμs (1 : ℝ)) a) • x := -begin - rw indicator_const_Lp_eq_to_span_singleton_comp_Lp hs hμs x, - have h_comp := condexp_L2_comp_continuous_linear_map ℝ 𝕜 hm (to_span_singleton ℝ x) - (indicator_const_Lp 2 hs hμs (1 : ℝ)), - rw ← Lp_meas_coe at h_comp, - refine h_comp.trans _, - exact (to_span_singleton ℝ x).coe_fn_comp_Lp _, -end - -lemma condexp_L2_indicator_eq_to_span_singleton_comp (hm : m ≤ m0) (hs : measurable_set s) - (hμs : μ s ≠ ∞) (x : E') : - (condexp_L2 𝕜 hm (indicator_const_Lp 2 hs hμs x) : α →₂[μ] E') - = (to_span_singleton ℝ x).comp_Lp (condexp_L2 ℝ hm (indicator_const_Lp 2 hs hμs (1 : ℝ))) := -begin - ext1, - rw ← Lp_meas_coe, - refine (condexp_L2_indicator_ae_eq_smul 𝕜 hm hs hμs x).trans _, - have h_comp := (to_span_singleton ℝ x).coe_fn_comp_Lp - (condexp_L2 ℝ hm (indicator_const_Lp 2 hs hμs (1 : ℝ)) : α →₂[μ] ℝ), - rw ← eventually_eq at h_comp, - refine eventually_eq.trans _ h_comp.symm, - refine eventually_of_forall (λ y, _), - refl, -end - -variables {𝕜} - -lemma set_lintegral_nnnorm_condexp_L2_indicator_le (hm : m ≤ m0) (hs : measurable_set s) - (hμs : μ s ≠ ∞) (x : E') {t : set α} (ht : measurable_set[m] t) (hμt : μ t ≠ ∞) : - ∫⁻ a in t, ‖condexp_L2 𝕜 hm (indicator_const_Lp 2 hs hμs x) a‖₊ ∂μ ≤ μ (s ∩ t) * ‖x‖₊ := -calc ∫⁻ a in t, ‖condexp_L2 𝕜 hm (indicator_const_Lp 2 hs hμs x) a‖₊ ∂μ - = ∫⁻ a in t, ‖(condexp_L2 ℝ hm (indicator_const_Lp 2 hs hμs (1 : ℝ)) a) • x‖₊ ∂μ : -set_lintegral_congr_fun (hm t ht) - ((condexp_L2_indicator_ae_eq_smul 𝕜 hm hs hμs x).mono (λ a ha hat, by rw ha)) -... = ∫⁻ a in t, ‖condexp_L2 ℝ hm (indicator_const_Lp 2 hs hμs (1 : ℝ)) a‖₊ ∂μ * ‖x‖₊ : -begin - simp_rw [nnnorm_smul, ennreal.coe_mul], - rw [lintegral_mul_const, Lp_meas_coe], - exact (Lp.strongly_measurable _).ennnorm -end -... ≤ μ (s ∩ t) * ‖x‖₊ : - mul_le_mul_right' (lintegral_nnnorm_condexp_L2_indicator_le_real hs hμs ht hμt) _ - -lemma lintegral_nnnorm_condexp_L2_indicator_le (hm : m ≤ m0) (hs : measurable_set s) - (hμs : μ s ≠ ∞) (x : E') [sigma_finite (μ.trim hm)] : - ∫⁻ a, ‖condexp_L2 𝕜 hm (indicator_const_Lp 2 hs hμs x) a‖₊ ∂μ ≤ μ s * ‖x‖₊ := -begin - refine lintegral_le_of_forall_fin_meas_le' hm (μ s * ‖x‖₊) _ (λ t ht hμt, _), - { rw Lp_meas_coe, - exact (Lp.ae_strongly_measurable _).ennnorm }, - refine (set_lintegral_nnnorm_condexp_L2_indicator_le hm hs hμs x ht hμt).trans _, - exact mul_le_mul_right' (measure_mono (set.inter_subset_left _ _)) _ -end - -/-- If the measure `μ.trim hm` is sigma-finite, then the conditional expectation of a measurable set -with finite measure is integrable. -/ -lemma integrable_condexp_L2_indicator (hm : m ≤ m0) [sigma_finite (μ.trim hm)] - (hs : measurable_set s) (hμs : μ s ≠ ∞) (x : E') : - integrable (condexp_L2 𝕜 hm (indicator_const_Lp 2 hs hμs x)) μ := -begin - refine integrable_of_forall_fin_meas_le' hm (μ s * ‖x‖₊) - (ennreal.mul_lt_top hμs ennreal.coe_ne_top) _ _, - { rw Lp_meas_coe, exact Lp.ae_strongly_measurable _, }, - { refine λ t ht hμt, (set_lintegral_nnnorm_condexp_L2_indicator_le hm hs hμs x ht hμt).trans _, - exact mul_le_mul_right' (measure_mono (set.inter_subset_left _ _)) _, }, -end - -end condexp_L2_indicator - -section condexp_ind_smul - -variables [normed_space ℝ G] {hm : m ≤ m0} - -/-- Conditional expectation of the indicator of a measurable set with finite measure, in L2. -/ -def condexp_ind_smul (hm : m ≤ m0) (hs : measurable_set s) (hμs : μ s ≠ ∞) (x : G) : Lp G 2 μ := -(to_span_singleton ℝ x).comp_LpL 2 μ (condexp_L2 ℝ hm (indicator_const_Lp 2 hs hμs (1 : ℝ))) - -lemma ae_strongly_measurable'_condexp_ind_smul - (hm : m ≤ m0) (hs : measurable_set s) (hμs : μ s ≠ ∞) (x : G) : - ae_strongly_measurable' m (condexp_ind_smul hm hs hμs x) μ := -begin - have h : ae_strongly_measurable' m (condexp_L2 ℝ hm (indicator_const_Lp 2 hs hμs (1 : ℝ))) μ, - from ae_strongly_measurable'_condexp_L2 _ _, - rw condexp_ind_smul, - suffices : ae_strongly_measurable' m - ((to_span_singleton ℝ x) ∘ (condexp_L2 ℝ hm (indicator_const_Lp 2 hs hμs (1 : ℝ)))) μ, - { refine ae_strongly_measurable'.congr this _, - refine eventually_eq.trans _ (coe_fn_comp_LpL _ _).symm, - rw Lp_meas_coe, }, - exact ae_strongly_measurable'.continuous_comp (to_span_singleton ℝ x).continuous h, -end - -lemma condexp_ind_smul_add (hs : measurable_set s) (hμs : μ s ≠ ∞) (x y : G) : - condexp_ind_smul hm hs hμs (x + y) - = condexp_ind_smul hm hs hμs x + condexp_ind_smul hm hs hμs y := -by { simp_rw [condexp_ind_smul], rw [to_span_singleton_add, add_comp_LpL, add_apply], } - -lemma condexp_ind_smul_smul (hs : measurable_set s) (hμs : μ s ≠ ∞) (c : ℝ) (x : G) : - condexp_ind_smul hm hs hμs (c • x) = c • condexp_ind_smul hm hs hμs x := -by { simp_rw [condexp_ind_smul], rw [to_span_singleton_smul, smul_comp_LpL, smul_apply], } - -lemma condexp_ind_smul_smul' [normed_space ℝ F] [smul_comm_class ℝ 𝕜 F] (hs : measurable_set s) - (hμs : μ s ≠ ∞) (c : 𝕜) (x : F) : - condexp_ind_smul hm hs hμs (c • x) = c • condexp_ind_smul hm hs hμs x := -by rw [condexp_ind_smul, condexp_ind_smul, to_span_singleton_smul', - (to_span_singleton ℝ x).smul_comp_LpL c, smul_apply] - -lemma condexp_ind_smul_ae_eq_smul (hm : m ≤ m0) (hs : measurable_set s) (hμs : μ s ≠ ∞) (x : G) : - condexp_ind_smul hm hs hμs x - =ᵐ[μ] λ a, (condexp_L2 ℝ hm (indicator_const_Lp 2 hs hμs (1 : ℝ)) a) • x := -(to_span_singleton ℝ x).coe_fn_comp_LpL _ - -lemma set_lintegral_nnnorm_condexp_ind_smul_le (hm : m ≤ m0) (hs : measurable_set s) - (hμs : μ s ≠ ∞) (x : G) {t : set α} (ht : measurable_set[m] t) (hμt : μ t ≠ ∞) : - ∫⁻ a in t, ‖condexp_ind_smul hm hs hμs x a‖₊ ∂μ ≤ μ (s ∩ t) * ‖x‖₊ := -calc ∫⁻ a in t, ‖condexp_ind_smul hm hs hμs x a‖₊ ∂μ - = ∫⁻ a in t, ‖condexp_L2 ℝ hm (indicator_const_Lp 2 hs hμs (1 : ℝ)) a • x‖₊ ∂μ : -set_lintegral_congr_fun (hm t ht) - ((condexp_ind_smul_ae_eq_smul hm hs hμs x).mono (λ a ha hat, by rw ha )) -... = ∫⁻ a in t, ‖condexp_L2 ℝ hm (indicator_const_Lp 2 hs hμs (1 : ℝ)) a‖₊ ∂μ * ‖x‖₊ : -begin - simp_rw [nnnorm_smul, ennreal.coe_mul], - rw [lintegral_mul_const, Lp_meas_coe], - exact (Lp.strongly_measurable _).ennnorm -end -... ≤ μ (s ∩ t) * ‖x‖₊ : - mul_le_mul_right' (lintegral_nnnorm_condexp_L2_indicator_le_real hs hμs ht hμt) _ - -lemma lintegral_nnnorm_condexp_ind_smul_le (hm : m ≤ m0) (hs : measurable_set s) - (hμs : μ s ≠ ∞) (x : G) [sigma_finite (μ.trim hm)] : - ∫⁻ a, ‖condexp_ind_smul hm hs hμs x a‖₊ ∂μ ≤ μ s * ‖x‖₊ := -begin - refine lintegral_le_of_forall_fin_meas_le' hm (μ s * ‖x‖₊) _ (λ t ht hμt, _), - { exact (Lp.ae_strongly_measurable _).ennnorm }, - refine (set_lintegral_nnnorm_condexp_ind_smul_le hm hs hμs x ht hμt).trans _, - exact mul_le_mul_right' (measure_mono (set.inter_subset_left _ _)) _ -end - -/-- If the measure `μ.trim hm` is sigma-finite, then the conditional expectation of a measurable set -with finite measure is integrable. -/ -lemma integrable_condexp_ind_smul (hm : m ≤ m0) [sigma_finite (μ.trim hm)] - (hs : measurable_set s) (hμs : μ s ≠ ∞) (x : G) : - integrable (condexp_ind_smul hm hs hμs x) μ := -begin - refine integrable_of_forall_fin_meas_le' hm (μ s * ‖x‖₊) - (ennreal.mul_lt_top hμs ennreal.coe_ne_top) _ _, - { exact Lp.ae_strongly_measurable _, }, - { refine λ t ht hμt, (set_lintegral_nnnorm_condexp_ind_smul_le hm hs hμs x ht hμt).trans _, - exact mul_le_mul_right' (measure_mono (set.inter_subset_left _ _)) _, }, -end - -lemma condexp_ind_smul_empty {x : G} : - condexp_ind_smul hm measurable_set.empty - ((@measure_empty _ _ μ).le.trans_lt ennreal.coe_lt_top).ne x = 0 := -begin - rw [condexp_ind_smul, indicator_const_empty], - simp only [coe_fn_coe_base, submodule.coe_zero, continuous_linear_map.map_zero], -end - -lemma set_integral_condexp_L2_indicator (hs : measurable_set[m] s) (ht : measurable_set t) - (hμs : μ s ≠ ∞) (hμt : μ t ≠ ∞) : - ∫ x in s, (condexp_L2 ℝ hm (indicator_const_Lp 2 ht hμt (1 : ℝ))) x ∂μ = (μ (t ∩ s)).to_real := -calc ∫ x in s, (condexp_L2 ℝ hm (indicator_const_Lp 2 ht hμt (1 : ℝ))) x ∂μ - = ∫ x in s, indicator_const_Lp 2 ht hμt (1 : ℝ) x ∂μ : - @integral_condexp_L2_eq - α _ ℝ _ _ _ _ _ _ _ _ _ hm (indicator_const_Lp 2 ht hμt (1 : ℝ)) hs hμs -... = (μ (t ∩ s)).to_real • 1 : set_integral_indicator_const_Lp (hm s hs) ht hμt (1 : ℝ) -... = (μ (t ∩ s)).to_real : by rw [smul_eq_mul, mul_one] - -lemma set_integral_condexp_ind_smul (hs : measurable_set[m] s) (ht : measurable_set t) - (hμs : μ s ≠ ∞) (hμt : μ t ≠ ∞) (x : G') : - ∫ a in s, (condexp_ind_smul hm ht hμt x) a ∂μ = (μ (t ∩ s)).to_real • x := -calc ∫ a in s, (condexp_ind_smul hm ht hμt x) a ∂μ - = (∫ a in s, (condexp_L2 ℝ hm (indicator_const_Lp 2 ht hμt (1 : ℝ)) a • x) ∂μ) : - set_integral_congr_ae (hm s hs) ((condexp_ind_smul_ae_eq_smul hm ht hμt x).mono (λ x hx hxs, hx)) -... = (∫ a in s, condexp_L2 ℝ hm (indicator_const_Lp 2 ht hμt (1 : ℝ)) a ∂μ) • x : - integral_smul_const _ x -... = (μ (t ∩ s)).to_real • x : - by rw set_integral_condexp_L2_indicator hs ht hμs hμt - -lemma condexp_L2_indicator_nonneg (hm : m ≤ m0) (hs : measurable_set s) (hμs : μ s ≠ ∞) - [sigma_finite (μ.trim hm)] : - 0 ≤ᵐ[μ] condexp_L2 ℝ hm (indicator_const_Lp 2 hs hμs (1 : ℝ)) := -begin - have h : ae_strongly_measurable' m (condexp_L2 ℝ hm (indicator_const_Lp 2 hs hμs (1 : ℝ))) μ, - from ae_strongly_measurable'_condexp_L2 _ _, - refine eventually_le.trans_eq _ h.ae_eq_mk.symm, - refine @ae_le_of_ae_le_trim _ _ _ _ _ _ hm _ _ _, - refine ae_nonneg_of_forall_set_integral_nonneg_of_sigma_finite _ _, - { intros t ht hμt, - refine @integrable.integrable_on _ _ m _ _ _ _ _, - refine integrable.trim hm _ _, - { rw integrable_congr h.ae_eq_mk.symm, - exact integrable_condexp_L2_indicator hm hs hμs _, }, - { exact h.strongly_measurable_mk, }, }, - { intros t ht hμt, - rw ← set_integral_trim hm h.strongly_measurable_mk ht, - have h_ae : ∀ᵐ x ∂μ, x ∈ t → h.mk _ x = condexp_L2 ℝ hm (indicator_const_Lp 2 hs hμs (1 : ℝ)) x, - { filter_upwards [h.ae_eq_mk] with x hx, - exact λ _, hx.symm, }, - rw [set_integral_congr_ae (hm t ht) h_ae, - set_integral_condexp_L2_indicator ht hs ((le_trim hm).trans_lt hμt).ne hμs], - exact ennreal.to_real_nonneg, }, -end - -lemma condexp_ind_smul_nonneg {E} [normed_lattice_add_comm_group E] [normed_space ℝ E] - [ordered_smul ℝ E] [sigma_finite (μ.trim hm)] - (hs : measurable_set s) (hμs : μ s ≠ ∞) (x : E) (hx : 0 ≤ x) : - 0 ≤ᵐ[μ] condexp_ind_smul hm hs hμs x := -begin - refine eventually_le.trans_eq _ (condexp_ind_smul_ae_eq_smul hm hs hμs x).symm, - filter_upwards [condexp_L2_indicator_nonneg hm hs hμs] with a ha, - exact smul_nonneg ha hx, -end - -end condexp_ind_smul - -end condexp_L2 - -section condexp_ind - -/-! ## Conditional expectation of an indicator as a continuous linear map. - -The goal of this section is to build -`condexp_ind (hm : m ≤ m0) (μ : measure α) (s : set s) : G →L[ℝ] α →₁[μ] G`, which -takes `x : G` to the conditional expectation of the indicator of the set `s` with value `x`, -seen as an element of `α →₁[μ] G`. --/ - -variables {m m0 : measurable_space α} {μ : measure α} {s t : set α} [normed_space ℝ G] - -section condexp_ind_L1_fin - -/-- Conditional expectation of the indicator of a measurable set with finite measure, -as a function in L1. -/ -def condexp_ind_L1_fin (hm : m ≤ m0) [sigma_finite (μ.trim hm)] (hs : measurable_set s) - (hμs : μ s ≠ ∞) (x : G) : α →₁[μ] G := -(integrable_condexp_ind_smul hm hs hμs x).to_L1 _ - -lemma condexp_ind_L1_fin_ae_eq_condexp_ind_smul (hm : m ≤ m0) [sigma_finite (μ.trim hm)] - (hs : measurable_set s) (hμs : μ s ≠ ∞) (x : G) : - condexp_ind_L1_fin hm hs hμs x =ᵐ[μ] condexp_ind_smul hm hs hμs x := -(integrable_condexp_ind_smul hm hs hμs x).coe_fn_to_L1 - -variables {hm : m ≤ m0} [sigma_finite (μ.trim hm)] - -lemma condexp_ind_L1_fin_add (hs : measurable_set s) (hμs : μ s ≠ ∞) (x y : G) : - condexp_ind_L1_fin hm hs hμs (x + y) - = condexp_ind_L1_fin hm hs hμs x + condexp_ind_L1_fin hm hs hμs y := -begin - ext1, - refine (mem_ℒp.coe_fn_to_Lp _).trans _, - refine eventually_eq.trans _ (Lp.coe_fn_add _ _).symm, - refine eventually_eq.trans _ - (eventually_eq.add (mem_ℒp.coe_fn_to_Lp _).symm (mem_ℒp.coe_fn_to_Lp _).symm), - rw condexp_ind_smul_add, - refine (Lp.coe_fn_add _ _).trans (eventually_of_forall (λ a, _)), - refl, -end - -lemma condexp_ind_L1_fin_smul (hs : measurable_set s) (hμs : μ s ≠ ∞) (c : ℝ) (x : G) : - condexp_ind_L1_fin hm hs hμs (c • x) = c • condexp_ind_L1_fin hm hs hμs x := -begin - ext1, - refine (mem_ℒp.coe_fn_to_Lp _).trans _, - refine eventually_eq.trans _ (Lp.coe_fn_smul _ _).symm, - rw condexp_ind_smul_smul hs hμs c x, - refine (Lp.coe_fn_smul _ _).trans _, - refine (condexp_ind_L1_fin_ae_eq_condexp_ind_smul hm hs hμs x).mono (λ y hy, _), - rw [pi.smul_apply, pi.smul_apply, hy], -end - -lemma condexp_ind_L1_fin_smul' [normed_space ℝ F] [smul_comm_class ℝ 𝕜 F] - (hs : measurable_set s) (hμs : μ s ≠ ∞) (c : 𝕜) (x : F) : - condexp_ind_L1_fin hm hs hμs (c • x) = c • condexp_ind_L1_fin hm hs hμs x := -begin - ext1, - refine (mem_ℒp.coe_fn_to_Lp _).trans _, - refine eventually_eq.trans _ (Lp.coe_fn_smul _ _).symm, - rw condexp_ind_smul_smul' hs hμs c x, - refine (Lp.coe_fn_smul _ _).trans _, - refine (condexp_ind_L1_fin_ae_eq_condexp_ind_smul hm hs hμs x).mono (λ y hy, _), - rw [pi.smul_apply, pi.smul_apply, hy], -end - -lemma norm_condexp_ind_L1_fin_le (hs : measurable_set s) (hμs : μ s ≠ ∞) (x : G) : - ‖condexp_ind_L1_fin hm hs hμs x‖ ≤ (μ s).to_real * ‖x‖ := -begin - have : 0 ≤ ∫ (a : α), ‖condexp_ind_L1_fin hm hs hμs x a‖ ∂μ, - from integral_nonneg (λ a, norm_nonneg _), - rw [L1.norm_eq_integral_norm, ← ennreal.to_real_of_real (norm_nonneg x), ← ennreal.to_real_mul, - ← ennreal.to_real_of_real this, ennreal.to_real_le_to_real ennreal.of_real_ne_top - (ennreal.mul_ne_top hμs ennreal.of_real_ne_top), - 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, ‖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, - rw hz, }, - rw [h_eq, of_real_norm_eq_coe_nnnorm], - exact lintegral_nnnorm_condexp_ind_smul_le hm hs hμs x, -end - -lemma condexp_ind_L1_fin_disjoint_union (hs : measurable_set s) (ht : measurable_set t) - (hμs : μ s ≠ ∞) (hμt : μ t ≠ ∞) (hst : s ∩ t = ∅) (x : G) : - condexp_ind_L1_fin hm (hs.union ht) ((measure_union_le s t).trans_lt - (lt_top_iff_ne_top.mpr (ennreal.add_ne_top.mpr ⟨hμs, hμt⟩))).ne x - = condexp_ind_L1_fin hm hs hμs x + condexp_ind_L1_fin hm ht hμt x := -begin - ext1, - have hμst := ((measure_union_le s t).trans_lt - (lt_top_iff_ne_top.mpr (ennreal.add_ne_top.mpr ⟨hμs, hμt⟩))).ne, - refine (condexp_ind_L1_fin_ae_eq_condexp_ind_smul hm (hs.union ht) hμst x).trans _, - refine eventually_eq.trans _ (Lp.coe_fn_add _ _).symm, - have hs_eq := condexp_ind_L1_fin_ae_eq_condexp_ind_smul hm hs hμs x, - have ht_eq := condexp_ind_L1_fin_ae_eq_condexp_ind_smul hm ht hμt x, - refine eventually_eq.trans _ (eventually_eq.add hs_eq.symm ht_eq.symm), - rw condexp_ind_smul, - rw indicator_const_Lp_disjoint_union hs ht hμs hμt hst (1 : ℝ), - rw (condexp_L2 ℝ hm).map_add, - push_cast, - rw ((to_span_singleton ℝ x).comp_LpL 2 μ).map_add, - refine (Lp.coe_fn_add _ _).trans _, - refine eventually_of_forall (λ y, _), - refl, -end - -end condexp_ind_L1_fin - -open_locale classical - -section condexp_ind_L1 - -/-- Conditional expectation of the indicator of a set, as a function in L1. Its value for sets -which are not both measurable and of finite measure is not used: we set it to 0. -/ -def condexp_ind_L1 {m m0 : measurable_space α} (hm : m ≤ m0) (μ : measure α) (s : set α) - [sigma_finite (μ.trim hm)] (x : G) : - α →₁[μ] G := -if hs : measurable_set s ∧ μ s ≠ ∞ then condexp_ind_L1_fin hm hs.1 hs.2 x else 0 - -variables {hm : m ≤ m0} [sigma_finite (μ.trim hm)] - -lemma condexp_ind_L1_of_measurable_set_of_measure_ne_top (hs : measurable_set s) (hμs : μ s ≠ ∞) - (x : G) : - condexp_ind_L1 hm μ s x = condexp_ind_L1_fin hm hs hμs x := -by simp only [condexp_ind_L1, and.intro hs hμs, dif_pos, ne.def, not_false_iff, and_self] - -lemma condexp_ind_L1_of_measure_eq_top (hμs : μ s = ∞) (x : G) : - condexp_ind_L1 hm μ s x = 0 := -by simp only [condexp_ind_L1, hμs, eq_self_iff_true, not_true, ne.def, dif_neg, not_false_iff, - and_false] - -lemma condexp_ind_L1_of_not_measurable_set (hs : ¬ measurable_set s) (x : G) : - condexp_ind_L1 hm μ s x = 0 := -by simp only [condexp_ind_L1, hs, dif_neg, not_false_iff, false_and] - -lemma condexp_ind_L1_add (x y : G) : - condexp_ind_L1 hm μ s (x + y) = condexp_ind_L1 hm μ s x + condexp_ind_L1 hm μ s y := -begin - by_cases hs : measurable_set s, - swap, {simp_rw condexp_ind_L1_of_not_measurable_set hs, rw zero_add, }, - by_cases hμs : μ s = ∞, - { simp_rw condexp_ind_L1_of_measure_eq_top hμs, rw zero_add, }, - { simp_rw condexp_ind_L1_of_measurable_set_of_measure_ne_top hs hμs, - exact condexp_ind_L1_fin_add hs hμs x y, }, -end - -lemma condexp_ind_L1_smul (c : ℝ) (x : G) : - condexp_ind_L1 hm μ s (c • x) = c • condexp_ind_L1 hm μ s x := -begin - by_cases hs : measurable_set s, - swap, {simp_rw condexp_ind_L1_of_not_measurable_set hs, rw smul_zero, }, - by_cases hμs : μ s = ∞, - { simp_rw condexp_ind_L1_of_measure_eq_top hμs, rw smul_zero, }, - { simp_rw condexp_ind_L1_of_measurable_set_of_measure_ne_top hs hμs, - exact condexp_ind_L1_fin_smul hs hμs c x, }, -end - -lemma condexp_ind_L1_smul' [normed_space ℝ F] [smul_comm_class ℝ 𝕜 F] (c : 𝕜) (x : F) : - condexp_ind_L1 hm μ s (c • x) = c • condexp_ind_L1 hm μ s x := -begin - by_cases hs : measurable_set s, - swap, {simp_rw condexp_ind_L1_of_not_measurable_set hs, rw smul_zero, }, - by_cases hμs : μ s = ∞, - { simp_rw condexp_ind_L1_of_measure_eq_top hμs, rw smul_zero, }, - { simp_rw condexp_ind_L1_of_measurable_set_of_measure_ne_top hs hμs, - exact condexp_ind_L1_fin_smul' hs hμs c x, }, -end - -lemma norm_condexp_ind_L1_le (x : G) : - ‖condexp_ind_L1 hm μ s x‖ ≤ (μ s).to_real * ‖x‖ := -begin - by_cases hs : measurable_set s, - swap, {simp_rw condexp_ind_L1_of_not_measurable_set hs, rw Lp.norm_zero, - exact mul_nonneg ennreal.to_real_nonneg (norm_nonneg _), }, - by_cases hμs : μ s = ∞, - { rw [condexp_ind_L1_of_measure_eq_top hμs x, Lp.norm_zero], - exact mul_nonneg ennreal.to_real_nonneg (norm_nonneg _), }, - { rw condexp_ind_L1_of_measurable_set_of_measure_ne_top hs hμs x, - exact norm_condexp_ind_L1_fin_le hs hμs x, }, -end - -lemma continuous_condexp_ind_L1 : continuous (λ x : G, condexp_ind_L1 hm μ s x) := -continuous_of_linear_of_bound condexp_ind_L1_add condexp_ind_L1_smul norm_condexp_ind_L1_le - -lemma condexp_ind_L1_disjoint_union (hs : measurable_set s) (ht : measurable_set t) - (hμs : μ s ≠ ∞) (hμt : μ t ≠ ∞) (hst : s ∩ t = ∅) (x : G) : - condexp_ind_L1 hm μ (s ∪ t) x = condexp_ind_L1 hm μ s x + condexp_ind_L1 hm μ t x := -begin - have hμst : μ (s ∪ t) ≠ ∞, from ((measure_union_le s t).trans_lt - (lt_top_iff_ne_top.mpr (ennreal.add_ne_top.mpr ⟨hμs, hμt⟩))).ne, - rw [condexp_ind_L1_of_measurable_set_of_measure_ne_top hs hμs x, - condexp_ind_L1_of_measurable_set_of_measure_ne_top ht hμt x, - condexp_ind_L1_of_measurable_set_of_measure_ne_top (hs.union ht) hμst x], - exact condexp_ind_L1_fin_disjoint_union hs ht hμs hμt hst x, -end - -end condexp_ind_L1 - -/-- Conditional expectation of the indicator of a set, as a linear map from `G` to L1. -/ -def condexp_ind {m m0 : measurable_space α} (hm : m ≤ m0) (μ : measure α) [sigma_finite (μ.trim hm)] - (s : set α) : G →L[ℝ] α →₁[μ] G := -{ to_fun := condexp_ind_L1 hm μ s, - map_add' := condexp_ind_L1_add, - map_smul' := condexp_ind_L1_smul, - cont := continuous_condexp_ind_L1, } - -lemma condexp_ind_ae_eq_condexp_ind_smul (hm : m ≤ m0) [sigma_finite (μ.trim hm)] - (hs : measurable_set s) (hμs : μ s ≠ ∞) (x : G) : - condexp_ind hm μ s x =ᵐ[μ] condexp_ind_smul hm hs hμs x := -begin - refine eventually_eq.trans _ (condexp_ind_L1_fin_ae_eq_condexp_ind_smul hm hs hμs x), - simp [condexp_ind, condexp_ind_L1, hs, hμs], -end - -variables {hm : m ≤ m0} [sigma_finite (μ.trim hm)] - -lemma ae_strongly_measurable'_condexp_ind (hs : measurable_set s) (hμs : μ s ≠ ∞) (x : G) : - ae_strongly_measurable' m (condexp_ind hm μ s x) μ := -ae_strongly_measurable'.congr (ae_strongly_measurable'_condexp_ind_smul hm hs hμs x) - (condexp_ind_ae_eq_condexp_ind_smul hm hs hμs x).symm - -@[simp] lemma condexp_ind_empty : condexp_ind hm μ ∅ = (0 : G →L[ℝ] α →₁[μ] G) := -begin - ext1, - ext1, - refine (condexp_ind_ae_eq_condexp_ind_smul hm measurable_set.empty (by simp) x).trans _, - rw condexp_ind_smul_empty, - refine (Lp.coe_fn_zero G 2 μ).trans _, - refine eventually_eq.trans _ (Lp.coe_fn_zero G 1 μ).symm, - refl, -end - -lemma condexp_ind_smul' [normed_space ℝ F] [smul_comm_class ℝ 𝕜 F] (c : 𝕜) (x : F) : - condexp_ind hm μ s (c • x) = c • condexp_ind hm μ s x := -condexp_ind_L1_smul' c x - -lemma norm_condexp_ind_apply_le (x : G) : ‖condexp_ind hm μ s x‖ ≤ (μ s).to_real * ‖x‖ := -norm_condexp_ind_L1_le x - -lemma norm_condexp_ind_le : ‖(condexp_ind hm μ s : G →L[ℝ] α →₁[μ] G)‖ ≤ (μ s).to_real := -continuous_linear_map.op_norm_le_bound _ ennreal.to_real_nonneg norm_condexp_ind_apply_le - -lemma condexp_ind_disjoint_union_apply (hs : measurable_set s) (ht : measurable_set t) - (hμs : μ s ≠ ∞) (hμt : μ t ≠ ∞) (hst : s ∩ t = ∅) (x : G) : - condexp_ind hm μ (s ∪ t) x = condexp_ind hm μ s x + condexp_ind hm μ t x := -condexp_ind_L1_disjoint_union hs ht hμs hμt hst x - -lemma condexp_ind_disjoint_union (hs : measurable_set s) (ht : measurable_set t) (hμs : μ s ≠ ∞) - (hμt : μ t ≠ ∞) (hst : s ∩ t = ∅) : - (condexp_ind hm μ (s ∪ t) : G →L[ℝ] α →₁[μ] G) = condexp_ind hm μ s + condexp_ind hm μ t := -by { ext1, push_cast, exact condexp_ind_disjoint_union_apply hs ht hμs hμt hst x, } - -variables (G) - -lemma dominated_fin_meas_additive_condexp_ind (hm : m ≤ m0) (μ : measure α) - [sigma_finite (μ.trim hm)] : - dominated_fin_meas_additive μ (condexp_ind hm μ : set α → G →L[ℝ] α →₁[μ] G) 1 := -⟨λ s t, condexp_ind_disjoint_union, λ s _ _, norm_condexp_ind_le.trans (one_mul _).symm.le⟩ - -variables {G} - -lemma set_integral_condexp_ind (hs : measurable_set[m] s) (ht : measurable_set t) (hμs : μ s ≠ ∞) - (hμt : μ t ≠ ∞) (x : G') : - ∫ a in s, condexp_ind hm μ t x a ∂μ = (μ (t ∩ s)).to_real • x := -calc -∫ a in s, condexp_ind hm μ t x a ∂μ = ∫ a in s, condexp_ind_smul hm ht hμt x a ∂μ : - set_integral_congr_ae (hm s hs) - ((condexp_ind_ae_eq_condexp_ind_smul hm ht hμt x).mono (λ x hx hxs, hx)) -... = (μ (t ∩ s)).to_real • x : set_integral_condexp_ind_smul hs ht hμs hμt x - -lemma condexp_ind_of_measurable (hs : measurable_set[m] s) (hμs : μ s ≠ ∞) (c : G) : - condexp_ind hm μ s c = indicator_const_Lp 1 (hm s hs) hμs c := -begin - ext1, - refine eventually_eq.trans _ indicator_const_Lp_coe_fn.symm, - refine (condexp_ind_ae_eq_condexp_ind_smul hm (hm s hs) hμs c).trans _, - refine (condexp_ind_smul_ae_eq_smul hm (hm s hs) hμs c).trans _, - rw [Lp_meas_coe, condexp_L2_indicator_of_measurable hm hs hμs (1 : ℝ)], - refine (@indicator_const_Lp_coe_fn α _ _ 2 μ _ s (hm s hs) hμs (1 : ℝ)).mono (λ x hx, _), - dsimp only, - rw hx, - by_cases hx_mem : x ∈ s; simp [hx_mem], -end - -lemma condexp_ind_nonneg {E} [normed_lattice_add_comm_group E] [normed_space ℝ E] [ordered_smul ℝ E] - (hs : measurable_set s) (hμs : μ s ≠ ∞) (x : E) (hx : 0 ≤ x) : - 0 ≤ condexp_ind hm μ s x := -begin - rw ← coe_fn_le, - refine eventually_le.trans_eq _ (condexp_ind_ae_eq_condexp_ind_smul hm hs hμs x).symm, - exact (coe_fn_zero E 1 μ).trans_le (condexp_ind_smul_nonneg hs hμs x hx), -end - -end condexp_ind - -section condexp_L1 - -variables {m m0 : measurable_space α} {μ : measure α} - {hm : m ≤ m0} [sigma_finite (μ.trim hm)] {f g : α → F'} {s : set α} - -/-- Conditional expectation of a function as a linear map from `α →₁[μ] F'` to itself. -/ -def condexp_L1_clm (hm : m ≤ m0) (μ : measure α) [sigma_finite (μ.trim hm)] : - (α →₁[μ] F') →L[ℝ] α →₁[μ] F' := -L1.set_to_L1 (dominated_fin_meas_additive_condexp_ind F' hm μ) - -lemma condexp_L1_clm_smul (c : 𝕜) (f : α →₁[μ] F') : - condexp_L1_clm hm μ (c • f) = c • condexp_L1_clm hm μ f := -L1.set_to_L1_smul (dominated_fin_meas_additive_condexp_ind F' hm μ) - (λ c s x, condexp_ind_smul' c x) c f - -lemma condexp_L1_clm_indicator_const_Lp (hs : measurable_set s) (hμs : μ s ≠ ∞) (x : F') : - (condexp_L1_clm hm μ) (indicator_const_Lp 1 hs hμs x) = condexp_ind hm μ s x := -L1.set_to_L1_indicator_const_Lp (dominated_fin_meas_additive_condexp_ind F' hm μ) hs hμs x - -lemma condexp_L1_clm_indicator_const (hs : measurable_set s) (hμs : μ s ≠ ∞) (x : F') : - (condexp_L1_clm hm μ) ↑(simple_func.indicator_const 1 hs hμs x) = condexp_ind hm μ s x := -by { rw Lp.simple_func.coe_indicator_const, exact condexp_L1_clm_indicator_const_Lp hs hμs x, } - -/-- Auxiliary lemma used in the proof of `set_integral_condexp_L1_clm`. -/ -lemma set_integral_condexp_L1_clm_of_measure_ne_top (f : α →₁[μ] F') (hs : measurable_set[m] s) - (hμs : μ s ≠ ∞) : - ∫ x in s, condexp_L1_clm hm μ f x ∂μ = ∫ x in s, f x ∂μ := -begin - refine Lp.induction ennreal.one_ne_top - (λ f : α →₁[μ] F', ∫ x in s, condexp_L1_clm hm μ f x ∂μ = ∫ x in s, f x ∂μ) - _ _ (is_closed_eq _ _) f, - { intros x t ht hμt, - simp_rw condexp_L1_clm_indicator_const ht hμt.ne x, - rw [Lp.simple_func.coe_indicator_const, set_integral_indicator_const_Lp (hm _ hs)], - exact set_integral_condexp_ind hs ht hμs hμt.ne x, }, - { intros f g hf_Lp hg_Lp hfg_disj hf hg, - simp_rw (condexp_L1_clm hm μ).map_add, - rw set_integral_congr_ae (hm s hs) ((Lp.coe_fn_add (condexp_L1_clm hm μ (hf_Lp.to_Lp f)) - (condexp_L1_clm hm μ (hg_Lp.to_Lp g))).mono (λ x hx hxs, hx)), - rw set_integral_congr_ae (hm s hs) ((Lp.coe_fn_add (hf_Lp.to_Lp f) (hg_Lp.to_Lp g)).mono - (λ x hx hxs, hx)), - simp_rw pi.add_apply, - rw [integral_add (L1.integrable_coe_fn _).integrable_on (L1.integrable_coe_fn _).integrable_on, - integral_add (L1.integrable_coe_fn _).integrable_on (L1.integrable_coe_fn _).integrable_on, - hf, hg], }, - { exact (continuous_set_integral s).comp (condexp_L1_clm hm μ).continuous, }, - { exact continuous_set_integral s, }, -end - -/-- The integral of the conditional expectation `condexp_L1_clm` over an `m`-measurable set is equal -to the integral of `f` on that set. See also `set_integral_condexp`, the similar statement for -`condexp`. -/ -lemma set_integral_condexp_L1_clm (f : α →₁[μ] F') (hs : measurable_set[m] s) : - ∫ x in s, condexp_L1_clm hm μ f x ∂μ = ∫ x in s, f x ∂μ := -begin - let S := spanning_sets (μ.trim hm), - have hS_meas : ∀ i, measurable_set[m] (S i) := measurable_spanning_sets (μ.trim hm), - have hS_meas0 : ∀ i, measurable_set (S i) := λ i, hm _ (hS_meas i), - have hs_eq : s = ⋃ i, S i ∩ s, - { simp_rw set.inter_comm, - rw [← set.inter_Union, (Union_spanning_sets (μ.trim hm)), set.inter_univ], }, - have hS_finite : ∀ i, μ (S i ∩ s) < ∞, - { refine λ i, (measure_mono (set.inter_subset_left _ _)).trans_lt _, - have hS_finite_trim := measure_spanning_sets_lt_top (μ.trim hm) i, - rwa trim_measurable_set_eq hm (hS_meas i) at hS_finite_trim, }, - have h_mono : monotone (λ i, (S i) ∩ s), - { intros i j hij x, - simp_rw set.mem_inter_iff, - exact λ h, ⟨monotone_spanning_sets (μ.trim hm) hij h.1, h.2⟩, }, - have h_eq_forall : (λ i, ∫ x in (S i) ∩ s, condexp_L1_clm hm μ f x ∂μ) - = λ i, ∫ x in (S i) ∩ s, f x ∂μ, - from funext (λ i, set_integral_condexp_L1_clm_of_measure_ne_top f - (@measurable_set.inter α m _ _ (hS_meas i) hs) (hS_finite i).ne), - have h_right : tendsto (λ i, ∫ x in (S i) ∩ s, f x ∂μ) at_top (𝓝 (∫ x in s, f x ∂μ)), - { have h := tendsto_set_integral_of_monotone (λ i, (hS_meas0 i).inter (hm s hs)) h_mono - (L1.integrable_coe_fn f).integrable_on, - rwa ← hs_eq at h, }, - have h_left : tendsto (λ i, ∫ x in (S i) ∩ s, condexp_L1_clm hm μ f x ∂μ) at_top - (𝓝 (∫ x in s, condexp_L1_clm hm μ f x ∂μ)), - { have h := tendsto_set_integral_of_monotone (λ i, (hS_meas0 i).inter (hm s hs)) - h_mono (L1.integrable_coe_fn (condexp_L1_clm hm μ f)).integrable_on, - rwa ← hs_eq at h, }, - rw h_eq_forall at h_left, - exact tendsto_nhds_unique h_left h_right, -end - -lemma ae_strongly_measurable'_condexp_L1_clm (f : α →₁[μ] F') : - ae_strongly_measurable' m (condexp_L1_clm hm μ f) μ := -begin - refine Lp.induction ennreal.one_ne_top - (λ f : α →₁[μ] F', ae_strongly_measurable' m (condexp_L1_clm hm μ f) μ) - _ _ _ f, - { intros c s hs hμs, - rw condexp_L1_clm_indicator_const hs hμs.ne c, - exact ae_strongly_measurable'_condexp_ind hs hμs.ne c, }, - { intros f g hf hg h_disj hfm hgm, - rw (condexp_L1_clm hm μ).map_add, - refine ae_strongly_measurable'.congr _ (coe_fn_add _ _).symm, - exact ae_strongly_measurable'.add hfm hgm, }, - { have : {f : Lp F' 1 μ | ae_strongly_measurable' m (condexp_L1_clm hm μ f) μ} - = (condexp_L1_clm hm μ) ⁻¹' {f | ae_strongly_measurable' m f μ}, - by refl, - rw this, - refine is_closed.preimage (condexp_L1_clm hm μ).continuous _, - exact is_closed_ae_strongly_measurable' hm, }, -end - -lemma condexp_L1_clm_Lp_meas (f : Lp_meas F' ℝ m 1 μ) : - condexp_L1_clm hm μ (f : α →₁[μ] F') = ↑f := -begin - let g := Lp_meas_to_Lp_trim_lie F' ℝ 1 μ hm f, - have hfg : f = (Lp_meas_to_Lp_trim_lie F' ℝ 1 μ hm).symm g, - by simp only [linear_isometry_equiv.symm_apply_apply], - rw hfg, - refine @Lp.induction α F' m _ 1 (μ.trim hm) _ ennreal.coe_ne_top - (λ g : α →₁[μ.trim hm] F', - condexp_L1_clm hm μ ((Lp_meas_to_Lp_trim_lie F' ℝ 1 μ hm).symm g : α →₁[μ] F') - = ↑((Lp_meas_to_Lp_trim_lie F' ℝ 1 μ hm).symm g)) _ _ _ g, - { intros c s hs hμs, - rw [Lp.simple_func.coe_indicator_const, Lp_meas_to_Lp_trim_lie_symm_indicator hs hμs.ne c, - condexp_L1_clm_indicator_const_Lp], - exact condexp_ind_of_measurable hs ((le_trim hm).trans_lt hμs).ne c, }, - { intros f g hf hg hfg_disj hf_eq hg_eq, - rw linear_isometry_equiv.map_add, - push_cast, - rw [map_add, hf_eq, hg_eq], }, - { refine is_closed_eq _ _, - { refine (condexp_L1_clm hm μ).continuous.comp (continuous_induced_dom.comp _), - exact linear_isometry_equiv.continuous _, }, - { refine continuous_induced_dom.comp _, - exact linear_isometry_equiv.continuous _, }, }, -end - -lemma condexp_L1_clm_of_ae_strongly_measurable' - (f : α →₁[μ] F') (hfm : ae_strongly_measurable' m f μ) : - condexp_L1_clm hm μ f = f := -condexp_L1_clm_Lp_meas (⟨f, hfm⟩ : Lp_meas F' ℝ m 1 μ) - -/-- Conditional expectation of a function, in L1. Its value is 0 if the function is not -integrable. The function-valued `condexp` should be used instead in most cases. -/ -def condexp_L1 (hm : m ≤ m0) (μ : measure α) [sigma_finite (μ.trim hm)] (f : α → F') : α →₁[μ] F' := -set_to_fun μ (condexp_ind hm μ) (dominated_fin_meas_additive_condexp_ind F' hm μ) f - -lemma condexp_L1_undef (hf : ¬ integrable f μ) : condexp_L1 hm μ f = 0 := -set_to_fun_undef (dominated_fin_meas_additive_condexp_ind F' hm μ) hf - -lemma condexp_L1_eq (hf : integrable f μ) : - condexp_L1 hm μ f = condexp_L1_clm hm μ (hf.to_L1 f) := -set_to_fun_eq (dominated_fin_meas_additive_condexp_ind F' hm μ) hf - -@[simp] lemma condexp_L1_zero : condexp_L1 hm μ (0 : α → F') = 0 := -set_to_fun_zero _ - -@[simp] lemma condexp_L1_measure_zero (hm : m ≤ m0) : condexp_L1 hm (0 : measure α) f = 0 := -set_to_fun_measure_zero _ rfl - -lemma ae_strongly_measurable'_condexp_L1 {f : α → F'} : - ae_strongly_measurable' m (condexp_L1 hm μ f) μ := -begin - by_cases hf : integrable f μ, - { rw condexp_L1_eq hf, - exact ae_strongly_measurable'_condexp_L1_clm _, }, - { rw condexp_L1_undef hf, - refine ae_strongly_measurable'.congr _ (coe_fn_zero _ _ _).symm, - exact strongly_measurable.ae_strongly_measurable' (@strongly_measurable_zero _ _ m _ _), }, -end - -lemma condexp_L1_congr_ae (hm : m ≤ m0) [sigma_finite (μ.trim hm)] (h : f =ᵐ[μ] g) : - condexp_L1 hm μ f = condexp_L1 hm μ g := -set_to_fun_congr_ae _ h - -lemma integrable_condexp_L1 (f : α → F') : integrable (condexp_L1 hm μ f) μ := -L1.integrable_coe_fn _ - -/-- The integral of the conditional expectation `condexp_L1` over an `m`-measurable set is equal to -the integral of `f` on that set. See also `set_integral_condexp`, the similar statement for -`condexp`. -/ -lemma set_integral_condexp_L1 (hf : integrable f μ) (hs : measurable_set[m] s) : - ∫ x in s, condexp_L1 hm μ f x ∂μ = ∫ x in s, f x ∂μ := -begin - simp_rw condexp_L1_eq hf, - rw set_integral_condexp_L1_clm (hf.to_L1 f) hs, - exact set_integral_congr_ae (hm s hs) ((hf.coe_fn_to_L1).mono (λ x hx hxs, hx)), -end - -lemma condexp_L1_add (hf : integrable f μ) (hg : integrable g μ) : - condexp_L1 hm μ (f + g) = condexp_L1 hm μ f + condexp_L1 hm μ g := -set_to_fun_add _ hf hg - -lemma condexp_L1_neg (f : α → F') : condexp_L1 hm μ (-f) = - condexp_L1 hm μ f := -set_to_fun_neg _ f - -lemma condexp_L1_smul (c : 𝕜) (f : α → F') : condexp_L1 hm μ (c • f) = c • condexp_L1 hm μ f := -set_to_fun_smul _ (λ c _ x, condexp_ind_smul' c x) c f - -lemma condexp_L1_sub (hf : integrable f μ) (hg : integrable g μ) : - condexp_L1 hm μ (f - g) = condexp_L1 hm μ f - condexp_L1 hm μ g := -set_to_fun_sub _ hf hg - -lemma condexp_L1_of_ae_strongly_measurable' - (hfm : ae_strongly_measurable' m f μ) (hfi : integrable f μ) : - condexp_L1 hm μ f =ᵐ[μ] f := -begin - rw condexp_L1_eq hfi, - refine eventually_eq.trans _ (integrable.coe_fn_to_L1 hfi), - rw condexp_L1_clm_of_ae_strongly_measurable', - exact ae_strongly_measurable'.congr hfm (integrable.coe_fn_to_L1 hfi).symm, -end - -lemma condexp_L1_mono {E} [normed_lattice_add_comm_group E] [complete_space E] [normed_space ℝ E] - [ordered_smul ℝ E] {f g : α → E} - (hf : integrable f μ) (hg : integrable g μ) (hfg : f ≤ᵐ[μ] g) : - condexp_L1 hm μ f ≤ᵐ[μ] condexp_L1 hm μ g := -begin - rw coe_fn_le, - have h_nonneg : ∀ s, measurable_set s → μ s < ∞ → ∀ x : E, 0 ≤ x → 0 ≤ condexp_ind hm μ s x, - from λ s hs hμs x hx, condexp_ind_nonneg hs hμs.ne x hx, - exact set_to_fun_mono (dominated_fin_meas_additive_condexp_ind E hm μ) h_nonneg hf hg hfg, -end - -end condexp_L1 - -section condexp - -/-! ### Conditional expectation of a function -/ open_locale classical @@ -1924,6 +84,7 @@ is true: - `μ` is not σ-finite with respect to `m`, - `f` is not integrable. -/ @[irreducible] +noncomputable def condexp (m : measurable_space α) {m0 : measurable_space α} (μ : measure α) (f : α → F') : α → F' := if hm : m ≤ m0 @@ -2280,6 +441,4 @@ begin exact tendsto_nhds_unique_of_eventually_eq hcond_gs hcond_fs (eventually_of_forall hn_eq), end -end condexp - end measure_theory diff --git a/src/measure_theory/function/conditional_expectation/condexp_L1.lean b/src/measure_theory/function/conditional_expectation/condexp_L1.lean new file mode 100644 index 0000000000000..1368135fae32e --- /dev/null +++ b/src/measure_theory/function/conditional_expectation/condexp_L1.lean @@ -0,0 +1,567 @@ +/- +Copyright (c) 2021 Rémy Degenne. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Rémy Degenne +-/ +import measure_theory.function.conditional_expectation.condexp_L2 + +/-! # Conditional expectation in L1 + +This file contains two more steps of the construction of the conditional expectation, which is +completed in `measure_theory.function.conditional_expectation.basic`. See that file for a +description of the full process. + +The contitional expectation of an `L²` function is defined in +`measure_theory.function.conditional_expectation.condexp_L2`. In this file, we perform two steps. +* Show that the conditional expectation of the indicator of a measurable set with finite measure + is integrable and define a map `set α → (E →L[ℝ] (α →₁[μ] E))` which to a set associates a linear + map. That linear map sends `x ∈ E` to the conditional expectation of the indicator of the set + with value `x`. +* Extend that map to `condexp_L1_clm : (α →₁[μ] E) →L[ℝ] (α →₁[μ] E)`. This is done using the same + construction as the Bochner integral (see the file `measure_theory/integral/set_to_L1`). + +## Main definitions + +* `condexp_L1`: Conditional expectation of a function as a linear map from `L1` to itself. + +-/ + +noncomputable theory +open topological_space measure_theory.Lp filter continuous_linear_map +open_locale nnreal ennreal topology big_operators measure_theory + +namespace measure_theory + +variables {α β F F' G G' 𝕜 : Type*} {p : ℝ≥0∞} + [is_R_or_C 𝕜] -- 𝕜 for ℝ or ℂ + -- F for a Lp submodule + [normed_add_comm_group F] [normed_space 𝕜 F] + -- F' for integrals on a Lp submodule + [normed_add_comm_group F'] [normed_space 𝕜 F'] [normed_space ℝ F'] [complete_space F'] + -- G for a Lp add_subgroup + [normed_add_comm_group G] + -- G' for integrals on a Lp add_subgroup + [normed_add_comm_group G'] [normed_space ℝ G'] [complete_space G'] + +section condexp_ind + +/-! ## Conditional expectation of an indicator as a continuous linear map. + +The goal of this section is to build +`condexp_ind (hm : m ≤ m0) (μ : measure α) (s : set s) : G →L[ℝ] α →₁[μ] G`, which +takes `x : G` to the conditional expectation of the indicator of the set `s` with value `x`, +seen as an element of `α →₁[μ] G`. +-/ + +variables {m m0 : measurable_space α} {μ : measure α} {s t : set α} [normed_space ℝ G] + +section condexp_ind_L1_fin + +/-- Conditional expectation of the indicator of a measurable set with finite measure, +as a function in L1. -/ +def condexp_ind_L1_fin (hm : m ≤ m0) [sigma_finite (μ.trim hm)] (hs : measurable_set s) + (hμs : μ s ≠ ∞) (x : G) : α →₁[μ] G := +(integrable_condexp_ind_smul hm hs hμs x).to_L1 _ + +lemma condexp_ind_L1_fin_ae_eq_condexp_ind_smul (hm : m ≤ m0) [sigma_finite (μ.trim hm)] + (hs : measurable_set s) (hμs : μ s ≠ ∞) (x : G) : + condexp_ind_L1_fin hm hs hμs x =ᵐ[μ] condexp_ind_smul hm hs hμs x := +(integrable_condexp_ind_smul hm hs hμs x).coe_fn_to_L1 + +variables {hm : m ≤ m0} [sigma_finite (μ.trim hm)] + +lemma condexp_ind_L1_fin_add (hs : measurable_set s) (hμs : μ s ≠ ∞) (x y : G) : + condexp_ind_L1_fin hm hs hμs (x + y) + = condexp_ind_L1_fin hm hs hμs x + condexp_ind_L1_fin hm hs hμs y := +begin + ext1, + refine (mem_ℒp.coe_fn_to_Lp _).trans _, + refine eventually_eq.trans _ (Lp.coe_fn_add _ _).symm, + refine eventually_eq.trans _ + (eventually_eq.add (mem_ℒp.coe_fn_to_Lp _).symm (mem_ℒp.coe_fn_to_Lp _).symm), + rw condexp_ind_smul_add, + refine (Lp.coe_fn_add _ _).trans (eventually_of_forall (λ a, _)), + refl, +end + +lemma condexp_ind_L1_fin_smul (hs : measurable_set s) (hμs : μ s ≠ ∞) (c : ℝ) (x : G) : + condexp_ind_L1_fin hm hs hμs (c • x) = c • condexp_ind_L1_fin hm hs hμs x := +begin + ext1, + refine (mem_ℒp.coe_fn_to_Lp _).trans _, + refine eventually_eq.trans _ (Lp.coe_fn_smul _ _).symm, + rw condexp_ind_smul_smul hs hμs c x, + refine (Lp.coe_fn_smul _ _).trans _, + refine (condexp_ind_L1_fin_ae_eq_condexp_ind_smul hm hs hμs x).mono (λ y hy, _), + rw [pi.smul_apply, pi.smul_apply, hy], +end + +lemma condexp_ind_L1_fin_smul' [normed_space ℝ F] [smul_comm_class ℝ 𝕜 F] + (hs : measurable_set s) (hμs : μ s ≠ ∞) (c : 𝕜) (x : F) : + condexp_ind_L1_fin hm hs hμs (c • x) = c • condexp_ind_L1_fin hm hs hμs x := +begin + ext1, + refine (mem_ℒp.coe_fn_to_Lp _).trans _, + refine eventually_eq.trans _ (Lp.coe_fn_smul _ _).symm, + rw condexp_ind_smul_smul' hs hμs c x, + refine (Lp.coe_fn_smul _ _).trans _, + refine (condexp_ind_L1_fin_ae_eq_condexp_ind_smul hm hs hμs x).mono (λ y hy, _), + rw [pi.smul_apply, pi.smul_apply, hy], +end + +lemma norm_condexp_ind_L1_fin_le (hs : measurable_set s) (hμs : μ s ≠ ∞) (x : G) : + ‖condexp_ind_L1_fin hm hs hμs x‖ ≤ (μ s).to_real * ‖x‖ := +begin + have : 0 ≤ ∫ (a : α), ‖condexp_ind_L1_fin hm hs hμs x a‖ ∂μ, + from integral_nonneg (λ a, norm_nonneg _), + rw [L1.norm_eq_integral_norm, ← ennreal.to_real_of_real (norm_nonneg x), ← ennreal.to_real_mul, + ← ennreal.to_real_of_real this, ennreal.to_real_le_to_real ennreal.of_real_ne_top + (ennreal.mul_ne_top hμs ennreal.of_real_ne_top), + 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, ‖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, + rw hz, }, + rw [h_eq, of_real_norm_eq_coe_nnnorm], + exact lintegral_nnnorm_condexp_ind_smul_le hm hs hμs x, +end + +lemma condexp_ind_L1_fin_disjoint_union (hs : measurable_set s) (ht : measurable_set t) + (hμs : μ s ≠ ∞) (hμt : μ t ≠ ∞) (hst : s ∩ t = ∅) (x : G) : + condexp_ind_L1_fin hm (hs.union ht) ((measure_union_le s t).trans_lt + (lt_top_iff_ne_top.mpr (ennreal.add_ne_top.mpr ⟨hμs, hμt⟩))).ne x + = condexp_ind_L1_fin hm hs hμs x + condexp_ind_L1_fin hm ht hμt x := +begin + ext1, + have hμst := ((measure_union_le s t).trans_lt + (lt_top_iff_ne_top.mpr (ennreal.add_ne_top.mpr ⟨hμs, hμt⟩))).ne, + refine (condexp_ind_L1_fin_ae_eq_condexp_ind_smul hm (hs.union ht) hμst x).trans _, + refine eventually_eq.trans _ (Lp.coe_fn_add _ _).symm, + have hs_eq := condexp_ind_L1_fin_ae_eq_condexp_ind_smul hm hs hμs x, + have ht_eq := condexp_ind_L1_fin_ae_eq_condexp_ind_smul hm ht hμt x, + refine eventually_eq.trans _ (eventually_eq.add hs_eq.symm ht_eq.symm), + rw condexp_ind_smul, + rw indicator_const_Lp_disjoint_union hs ht hμs hμt hst (1 : ℝ), + rw (condexp_L2 ℝ hm).map_add, + push_cast, + rw ((to_span_singleton ℝ x).comp_LpL 2 μ).map_add, + refine (Lp.coe_fn_add _ _).trans _, + refine eventually_of_forall (λ y, _), + refl, +end + +end condexp_ind_L1_fin + +open_locale classical + +section condexp_ind_L1 + +/-- Conditional expectation of the indicator of a set, as a function in L1. Its value for sets +which are not both measurable and of finite measure is not used: we set it to 0. -/ +def condexp_ind_L1 {m m0 : measurable_space α} (hm : m ≤ m0) (μ : measure α) (s : set α) + [sigma_finite (μ.trim hm)] (x : G) : + α →₁[μ] G := +if hs : measurable_set s ∧ μ s ≠ ∞ then condexp_ind_L1_fin hm hs.1 hs.2 x else 0 + +variables {hm : m ≤ m0} [sigma_finite (μ.trim hm)] + +lemma condexp_ind_L1_of_measurable_set_of_measure_ne_top (hs : measurable_set s) (hμs : μ s ≠ ∞) + (x : G) : + condexp_ind_L1 hm μ s x = condexp_ind_L1_fin hm hs hμs x := +by simp only [condexp_ind_L1, and.intro hs hμs, dif_pos, ne.def, not_false_iff, and_self] + +lemma condexp_ind_L1_of_measure_eq_top (hμs : μ s = ∞) (x : G) : + condexp_ind_L1 hm μ s x = 0 := +by simp only [condexp_ind_L1, hμs, eq_self_iff_true, not_true, ne.def, dif_neg, not_false_iff, + and_false] + +lemma condexp_ind_L1_of_not_measurable_set (hs : ¬ measurable_set s) (x : G) : + condexp_ind_L1 hm μ s x = 0 := +by simp only [condexp_ind_L1, hs, dif_neg, not_false_iff, false_and] + +lemma condexp_ind_L1_add (x y : G) : + condexp_ind_L1 hm μ s (x + y) = condexp_ind_L1 hm μ s x + condexp_ind_L1 hm μ s y := +begin + by_cases hs : measurable_set s, + swap, {simp_rw condexp_ind_L1_of_not_measurable_set hs, rw zero_add, }, + by_cases hμs : μ s = ∞, + { simp_rw condexp_ind_L1_of_measure_eq_top hμs, rw zero_add, }, + { simp_rw condexp_ind_L1_of_measurable_set_of_measure_ne_top hs hμs, + exact condexp_ind_L1_fin_add hs hμs x y, }, +end + +lemma condexp_ind_L1_smul (c : ℝ) (x : G) : + condexp_ind_L1 hm μ s (c • x) = c • condexp_ind_L1 hm μ s x := +begin + by_cases hs : measurable_set s, + swap, {simp_rw condexp_ind_L1_of_not_measurable_set hs, rw smul_zero, }, + by_cases hμs : μ s = ∞, + { simp_rw condexp_ind_L1_of_measure_eq_top hμs, rw smul_zero, }, + { simp_rw condexp_ind_L1_of_measurable_set_of_measure_ne_top hs hμs, + exact condexp_ind_L1_fin_smul hs hμs c x, }, +end + +lemma condexp_ind_L1_smul' [normed_space ℝ F] [smul_comm_class ℝ 𝕜 F] (c : 𝕜) (x : F) : + condexp_ind_L1 hm μ s (c • x) = c • condexp_ind_L1 hm μ s x := +begin + by_cases hs : measurable_set s, + swap, {simp_rw condexp_ind_L1_of_not_measurable_set hs, rw smul_zero, }, + by_cases hμs : μ s = ∞, + { simp_rw condexp_ind_L1_of_measure_eq_top hμs, rw smul_zero, }, + { simp_rw condexp_ind_L1_of_measurable_set_of_measure_ne_top hs hμs, + exact condexp_ind_L1_fin_smul' hs hμs c x, }, +end + +lemma norm_condexp_ind_L1_le (x : G) : + ‖condexp_ind_L1 hm μ s x‖ ≤ (μ s).to_real * ‖x‖ := +begin + by_cases hs : measurable_set s, + swap, {simp_rw condexp_ind_L1_of_not_measurable_set hs, rw Lp.norm_zero, + exact mul_nonneg ennreal.to_real_nonneg (norm_nonneg _), }, + by_cases hμs : μ s = ∞, + { rw [condexp_ind_L1_of_measure_eq_top hμs x, Lp.norm_zero], + exact mul_nonneg ennreal.to_real_nonneg (norm_nonneg _), }, + { rw condexp_ind_L1_of_measurable_set_of_measure_ne_top hs hμs x, + exact norm_condexp_ind_L1_fin_le hs hμs x, }, +end + +lemma continuous_condexp_ind_L1 : continuous (λ x : G, condexp_ind_L1 hm μ s x) := +continuous_of_linear_of_bound condexp_ind_L1_add condexp_ind_L1_smul norm_condexp_ind_L1_le + +lemma condexp_ind_L1_disjoint_union (hs : measurable_set s) (ht : measurable_set t) + (hμs : μ s ≠ ∞) (hμt : μ t ≠ ∞) (hst : s ∩ t = ∅) (x : G) : + condexp_ind_L1 hm μ (s ∪ t) x = condexp_ind_L1 hm μ s x + condexp_ind_L1 hm μ t x := +begin + have hμst : μ (s ∪ t) ≠ ∞, from ((measure_union_le s t).trans_lt + (lt_top_iff_ne_top.mpr (ennreal.add_ne_top.mpr ⟨hμs, hμt⟩))).ne, + rw [condexp_ind_L1_of_measurable_set_of_measure_ne_top hs hμs x, + condexp_ind_L1_of_measurable_set_of_measure_ne_top ht hμt x, + condexp_ind_L1_of_measurable_set_of_measure_ne_top (hs.union ht) hμst x], + exact condexp_ind_L1_fin_disjoint_union hs ht hμs hμt hst x, +end + +end condexp_ind_L1 + +/-- Conditional expectation of the indicator of a set, as a linear map from `G` to L1. -/ +def condexp_ind {m m0 : measurable_space α} (hm : m ≤ m0) (μ : measure α) [sigma_finite (μ.trim hm)] + (s : set α) : G →L[ℝ] α →₁[μ] G := +{ to_fun := condexp_ind_L1 hm μ s, + map_add' := condexp_ind_L1_add, + map_smul' := condexp_ind_L1_smul, + cont := continuous_condexp_ind_L1, } + +lemma condexp_ind_ae_eq_condexp_ind_smul (hm : m ≤ m0) [sigma_finite (μ.trim hm)] + (hs : measurable_set s) (hμs : μ s ≠ ∞) (x : G) : + condexp_ind hm μ s x =ᵐ[μ] condexp_ind_smul hm hs hμs x := +begin + refine eventually_eq.trans _ (condexp_ind_L1_fin_ae_eq_condexp_ind_smul hm hs hμs x), + simp [condexp_ind, condexp_ind_L1, hs, hμs], +end + +variables {hm : m ≤ m0} [sigma_finite (μ.trim hm)] + +lemma ae_strongly_measurable'_condexp_ind (hs : measurable_set s) (hμs : μ s ≠ ∞) (x : G) : + ae_strongly_measurable' m (condexp_ind hm μ s x) μ := +ae_strongly_measurable'.congr (ae_strongly_measurable'_condexp_ind_smul hm hs hμs x) + (condexp_ind_ae_eq_condexp_ind_smul hm hs hμs x).symm + +@[simp] lemma condexp_ind_empty : condexp_ind hm μ ∅ = (0 : G →L[ℝ] α →₁[μ] G) := +begin + ext1, + ext1, + refine (condexp_ind_ae_eq_condexp_ind_smul hm measurable_set.empty (by simp) x).trans _, + rw condexp_ind_smul_empty, + refine (Lp.coe_fn_zero G 2 μ).trans _, + refine eventually_eq.trans _ (Lp.coe_fn_zero G 1 μ).symm, + refl, +end + +lemma condexp_ind_smul' [normed_space ℝ F] [smul_comm_class ℝ 𝕜 F] (c : 𝕜) (x : F) : + condexp_ind hm μ s (c • x) = c • condexp_ind hm μ s x := +condexp_ind_L1_smul' c x + +lemma norm_condexp_ind_apply_le (x : G) : ‖condexp_ind hm μ s x‖ ≤ (μ s).to_real * ‖x‖ := +norm_condexp_ind_L1_le x + +lemma norm_condexp_ind_le : ‖(condexp_ind hm μ s : G →L[ℝ] α →₁[μ] G)‖ ≤ (μ s).to_real := +continuous_linear_map.op_norm_le_bound _ ennreal.to_real_nonneg norm_condexp_ind_apply_le + +lemma condexp_ind_disjoint_union_apply (hs : measurable_set s) (ht : measurable_set t) + (hμs : μ s ≠ ∞) (hμt : μ t ≠ ∞) (hst : s ∩ t = ∅) (x : G) : + condexp_ind hm μ (s ∪ t) x = condexp_ind hm μ s x + condexp_ind hm μ t x := +condexp_ind_L1_disjoint_union hs ht hμs hμt hst x + +lemma condexp_ind_disjoint_union (hs : measurable_set s) (ht : measurable_set t) (hμs : μ s ≠ ∞) + (hμt : μ t ≠ ∞) (hst : s ∩ t = ∅) : + (condexp_ind hm μ (s ∪ t) : G →L[ℝ] α →₁[μ] G) = condexp_ind hm μ s + condexp_ind hm μ t := +by { ext1, push_cast, exact condexp_ind_disjoint_union_apply hs ht hμs hμt hst x, } + +variables (G) + +lemma dominated_fin_meas_additive_condexp_ind (hm : m ≤ m0) (μ : measure α) + [sigma_finite (μ.trim hm)] : + dominated_fin_meas_additive μ (condexp_ind hm μ : set α → G →L[ℝ] α →₁[μ] G) 1 := +⟨λ s t, condexp_ind_disjoint_union, λ s _ _, norm_condexp_ind_le.trans (one_mul _).symm.le⟩ + +variables {G} + +lemma set_integral_condexp_ind (hs : measurable_set[m] s) (ht : measurable_set t) (hμs : μ s ≠ ∞) + (hμt : μ t ≠ ∞) (x : G') : + ∫ a in s, condexp_ind hm μ t x a ∂μ = (μ (t ∩ s)).to_real • x := +calc +∫ a in s, condexp_ind hm μ t x a ∂μ = ∫ a in s, condexp_ind_smul hm ht hμt x a ∂μ : + set_integral_congr_ae (hm s hs) + ((condexp_ind_ae_eq_condexp_ind_smul hm ht hμt x).mono (λ x hx hxs, hx)) +... = (μ (t ∩ s)).to_real • x : set_integral_condexp_ind_smul hs ht hμs hμt x + +lemma condexp_ind_of_measurable (hs : measurable_set[m] s) (hμs : μ s ≠ ∞) (c : G) : + condexp_ind hm μ s c = indicator_const_Lp 1 (hm s hs) hμs c := +begin + ext1, + refine eventually_eq.trans _ indicator_const_Lp_coe_fn.symm, + refine (condexp_ind_ae_eq_condexp_ind_smul hm (hm s hs) hμs c).trans _, + refine (condexp_ind_smul_ae_eq_smul hm (hm s hs) hμs c).trans _, + rw [Lp_meas_coe, condexp_L2_indicator_of_measurable hm hs hμs (1 : ℝ)], + refine (@indicator_const_Lp_coe_fn α _ _ 2 μ _ s (hm s hs) hμs (1 : ℝ)).mono (λ x hx, _), + dsimp only, + rw hx, + by_cases hx_mem : x ∈ s; simp [hx_mem], +end + +lemma condexp_ind_nonneg {E} [normed_lattice_add_comm_group E] [normed_space ℝ E] [ordered_smul ℝ E] + (hs : measurable_set s) (hμs : μ s ≠ ∞) (x : E) (hx : 0 ≤ x) : + 0 ≤ condexp_ind hm μ s x := +begin + rw ← coe_fn_le, + refine eventually_le.trans_eq _ (condexp_ind_ae_eq_condexp_ind_smul hm hs hμs x).symm, + exact (coe_fn_zero E 1 μ).trans_le (condexp_ind_smul_nonneg hs hμs x hx), +end + +end condexp_ind + +section condexp_L1 + +variables {m m0 : measurable_space α} {μ : measure α} + {hm : m ≤ m0} [sigma_finite (μ.trim hm)] {f g : α → F'} {s : set α} + +/-- Conditional expectation of a function as a linear map from `α →₁[μ] F'` to itself. -/ +def condexp_L1_clm (hm : m ≤ m0) (μ : measure α) [sigma_finite (μ.trim hm)] : + (α →₁[μ] F') →L[ℝ] α →₁[μ] F' := +L1.set_to_L1 (dominated_fin_meas_additive_condexp_ind F' hm μ) + +lemma condexp_L1_clm_smul (c : 𝕜) (f : α →₁[μ] F') : + condexp_L1_clm hm μ (c • f) = c • condexp_L1_clm hm μ f := +L1.set_to_L1_smul (dominated_fin_meas_additive_condexp_ind F' hm μ) + (λ c s x, condexp_ind_smul' c x) c f + +lemma condexp_L1_clm_indicator_const_Lp (hs : measurable_set s) (hμs : μ s ≠ ∞) (x : F') : + (condexp_L1_clm hm μ) (indicator_const_Lp 1 hs hμs x) = condexp_ind hm μ s x := +L1.set_to_L1_indicator_const_Lp (dominated_fin_meas_additive_condexp_ind F' hm μ) hs hμs x + +lemma condexp_L1_clm_indicator_const (hs : measurable_set s) (hμs : μ s ≠ ∞) (x : F') : + (condexp_L1_clm hm μ) ↑(simple_func.indicator_const 1 hs hμs x) = condexp_ind hm μ s x := +by { rw Lp.simple_func.coe_indicator_const, exact condexp_L1_clm_indicator_const_Lp hs hμs x, } + +/-- Auxiliary lemma used in the proof of `set_integral_condexp_L1_clm`. -/ +lemma set_integral_condexp_L1_clm_of_measure_ne_top (f : α →₁[μ] F') (hs : measurable_set[m] s) + (hμs : μ s ≠ ∞) : + ∫ x in s, condexp_L1_clm hm μ f x ∂μ = ∫ x in s, f x ∂μ := +begin + refine Lp.induction ennreal.one_ne_top + (λ f : α →₁[μ] F', ∫ x in s, condexp_L1_clm hm μ f x ∂μ = ∫ x in s, f x ∂μ) + _ _ (is_closed_eq _ _) f, + { intros x t ht hμt, + simp_rw condexp_L1_clm_indicator_const ht hμt.ne x, + rw [Lp.simple_func.coe_indicator_const, set_integral_indicator_const_Lp (hm _ hs)], + exact set_integral_condexp_ind hs ht hμs hμt.ne x, }, + { intros f g hf_Lp hg_Lp hfg_disj hf hg, + simp_rw (condexp_L1_clm hm μ).map_add, + rw set_integral_congr_ae (hm s hs) ((Lp.coe_fn_add (condexp_L1_clm hm μ (hf_Lp.to_Lp f)) + (condexp_L1_clm hm μ (hg_Lp.to_Lp g))).mono (λ x hx hxs, hx)), + rw set_integral_congr_ae (hm s hs) ((Lp.coe_fn_add (hf_Lp.to_Lp f) (hg_Lp.to_Lp g)).mono + (λ x hx hxs, hx)), + simp_rw pi.add_apply, + rw [integral_add (L1.integrable_coe_fn _).integrable_on (L1.integrable_coe_fn _).integrable_on, + integral_add (L1.integrable_coe_fn _).integrable_on (L1.integrable_coe_fn _).integrable_on, + hf, hg], }, + { exact (continuous_set_integral s).comp (condexp_L1_clm hm μ).continuous, }, + { exact continuous_set_integral s, }, +end + +/-- The integral of the conditional expectation `condexp_L1_clm` over an `m`-measurable set is equal +to the integral of `f` on that set. See also `set_integral_condexp`, the similar statement for +`condexp`. -/ +lemma set_integral_condexp_L1_clm (f : α →₁[μ] F') (hs : measurable_set[m] s) : + ∫ x in s, condexp_L1_clm hm μ f x ∂μ = ∫ x in s, f x ∂μ := +begin + let S := spanning_sets (μ.trim hm), + have hS_meas : ∀ i, measurable_set[m] (S i) := measurable_spanning_sets (μ.trim hm), + have hS_meas0 : ∀ i, measurable_set (S i) := λ i, hm _ (hS_meas i), + have hs_eq : s = ⋃ i, S i ∩ s, + { simp_rw set.inter_comm, + rw [← set.inter_Union, (Union_spanning_sets (μ.trim hm)), set.inter_univ], }, + have hS_finite : ∀ i, μ (S i ∩ s) < ∞, + { refine λ i, (measure_mono (set.inter_subset_left _ _)).trans_lt _, + have hS_finite_trim := measure_spanning_sets_lt_top (μ.trim hm) i, + rwa trim_measurable_set_eq hm (hS_meas i) at hS_finite_trim, }, + have h_mono : monotone (λ i, (S i) ∩ s), + { intros i j hij x, + simp_rw set.mem_inter_iff, + exact λ h, ⟨monotone_spanning_sets (μ.trim hm) hij h.1, h.2⟩, }, + have h_eq_forall : (λ i, ∫ x in (S i) ∩ s, condexp_L1_clm hm μ f x ∂μ) + = λ i, ∫ x in (S i) ∩ s, f x ∂μ, + from funext (λ i, set_integral_condexp_L1_clm_of_measure_ne_top f + (@measurable_set.inter α m _ _ (hS_meas i) hs) (hS_finite i).ne), + have h_right : tendsto (λ i, ∫ x in (S i) ∩ s, f x ∂μ) at_top (𝓝 (∫ x in s, f x ∂μ)), + { have h := tendsto_set_integral_of_monotone (λ i, (hS_meas0 i).inter (hm s hs)) h_mono + (L1.integrable_coe_fn f).integrable_on, + rwa ← hs_eq at h, }, + have h_left : tendsto (λ i, ∫ x in (S i) ∩ s, condexp_L1_clm hm μ f x ∂μ) at_top + (𝓝 (∫ x in s, condexp_L1_clm hm μ f x ∂μ)), + { have h := tendsto_set_integral_of_monotone (λ i, (hS_meas0 i).inter (hm s hs)) + h_mono (L1.integrable_coe_fn (condexp_L1_clm hm μ f)).integrable_on, + rwa ← hs_eq at h, }, + rw h_eq_forall at h_left, + exact tendsto_nhds_unique h_left h_right, +end + +lemma ae_strongly_measurable'_condexp_L1_clm (f : α →₁[μ] F') : + ae_strongly_measurable' m (condexp_L1_clm hm μ f) μ := +begin + refine Lp.induction ennreal.one_ne_top + (λ f : α →₁[μ] F', ae_strongly_measurable' m (condexp_L1_clm hm μ f) μ) + _ _ _ f, + { intros c s hs hμs, + rw condexp_L1_clm_indicator_const hs hμs.ne c, + exact ae_strongly_measurable'_condexp_ind hs hμs.ne c, }, + { intros f g hf hg h_disj hfm hgm, + rw (condexp_L1_clm hm μ).map_add, + refine ae_strongly_measurable'.congr _ (coe_fn_add _ _).symm, + exact ae_strongly_measurable'.add hfm hgm, }, + { have : {f : Lp F' 1 μ | ae_strongly_measurable' m (condexp_L1_clm hm μ f) μ} + = (condexp_L1_clm hm μ) ⁻¹' {f | ae_strongly_measurable' m f μ}, + by refl, + rw this, + refine is_closed.preimage (condexp_L1_clm hm μ).continuous _, + exact is_closed_ae_strongly_measurable' hm, }, +end + +lemma condexp_L1_clm_Lp_meas (f : Lp_meas F' ℝ m 1 μ) : + condexp_L1_clm hm μ (f : α →₁[μ] F') = ↑f := +begin + let g := Lp_meas_to_Lp_trim_lie F' ℝ 1 μ hm f, + have hfg : f = (Lp_meas_to_Lp_trim_lie F' ℝ 1 μ hm).symm g, + by simp only [linear_isometry_equiv.symm_apply_apply], + rw hfg, + refine @Lp.induction α F' m _ 1 (μ.trim hm) _ ennreal.coe_ne_top + (λ g : α →₁[μ.trim hm] F', + condexp_L1_clm hm μ ((Lp_meas_to_Lp_trim_lie F' ℝ 1 μ hm).symm g : α →₁[μ] F') + = ↑((Lp_meas_to_Lp_trim_lie F' ℝ 1 μ hm).symm g)) _ _ _ g, + { intros c s hs hμs, + rw [Lp.simple_func.coe_indicator_const, Lp_meas_to_Lp_trim_lie_symm_indicator hs hμs.ne c, + condexp_L1_clm_indicator_const_Lp], + exact condexp_ind_of_measurable hs ((le_trim hm).trans_lt hμs).ne c, }, + { intros f g hf hg hfg_disj hf_eq hg_eq, + rw linear_isometry_equiv.map_add, + push_cast, + rw [map_add, hf_eq, hg_eq], }, + { refine is_closed_eq _ _, + { refine (condexp_L1_clm hm μ).continuous.comp (continuous_induced_dom.comp _), + exact linear_isometry_equiv.continuous _, }, + { refine continuous_induced_dom.comp _, + exact linear_isometry_equiv.continuous _, }, }, +end + +lemma condexp_L1_clm_of_ae_strongly_measurable' + (f : α →₁[μ] F') (hfm : ae_strongly_measurable' m f μ) : + condexp_L1_clm hm μ f = f := +condexp_L1_clm_Lp_meas (⟨f, hfm⟩ : Lp_meas F' ℝ m 1 μ) + +/-- Conditional expectation of a function, in L1. Its value is 0 if the function is not +integrable. The function-valued `condexp` should be used instead in most cases. -/ +def condexp_L1 (hm : m ≤ m0) (μ : measure α) [sigma_finite (μ.trim hm)] (f : α → F') : α →₁[μ] F' := +set_to_fun μ (condexp_ind hm μ) (dominated_fin_meas_additive_condexp_ind F' hm μ) f + +lemma condexp_L1_undef (hf : ¬ integrable f μ) : condexp_L1 hm μ f = 0 := +set_to_fun_undef (dominated_fin_meas_additive_condexp_ind F' hm μ) hf + +lemma condexp_L1_eq (hf : integrable f μ) : + condexp_L1 hm μ f = condexp_L1_clm hm μ (hf.to_L1 f) := +set_to_fun_eq (dominated_fin_meas_additive_condexp_ind F' hm μ) hf + +@[simp] lemma condexp_L1_zero : condexp_L1 hm μ (0 : α → F') = 0 := +set_to_fun_zero _ + +@[simp] lemma condexp_L1_measure_zero (hm : m ≤ m0) : condexp_L1 hm (0 : measure α) f = 0 := +set_to_fun_measure_zero _ rfl + +lemma ae_strongly_measurable'_condexp_L1 {f : α → F'} : + ae_strongly_measurable' m (condexp_L1 hm μ f) μ := +begin + by_cases hf : integrable f μ, + { rw condexp_L1_eq hf, + exact ae_strongly_measurable'_condexp_L1_clm _, }, + { rw condexp_L1_undef hf, + refine ae_strongly_measurable'.congr _ (coe_fn_zero _ _ _).symm, + exact strongly_measurable.ae_strongly_measurable' (@strongly_measurable_zero _ _ m _ _), }, +end + +lemma condexp_L1_congr_ae (hm : m ≤ m0) [sigma_finite (μ.trim hm)] (h : f =ᵐ[μ] g) : + condexp_L1 hm μ f = condexp_L1 hm μ g := +set_to_fun_congr_ae _ h + +lemma integrable_condexp_L1 (f : α → F') : integrable (condexp_L1 hm μ f) μ := +L1.integrable_coe_fn _ + +/-- The integral of the conditional expectation `condexp_L1` over an `m`-measurable set is equal to +the integral of `f` on that set. See also `set_integral_condexp`, the similar statement for +`condexp`. -/ +lemma set_integral_condexp_L1 (hf : integrable f μ) (hs : measurable_set[m] s) : + ∫ x in s, condexp_L1 hm μ f x ∂μ = ∫ x in s, f x ∂μ := +begin + simp_rw condexp_L1_eq hf, + rw set_integral_condexp_L1_clm (hf.to_L1 f) hs, + exact set_integral_congr_ae (hm s hs) ((hf.coe_fn_to_L1).mono (λ x hx hxs, hx)), +end + +lemma condexp_L1_add (hf : integrable f μ) (hg : integrable g μ) : + condexp_L1 hm μ (f + g) = condexp_L1 hm μ f + condexp_L1 hm μ g := +set_to_fun_add _ hf hg + +lemma condexp_L1_neg (f : α → F') : condexp_L1 hm μ (-f) = - condexp_L1 hm μ f := +set_to_fun_neg _ f + +lemma condexp_L1_smul (c : 𝕜) (f : α → F') : condexp_L1 hm μ (c • f) = c • condexp_L1 hm μ f := +set_to_fun_smul _ (λ c _ x, condexp_ind_smul' c x) c f + +lemma condexp_L1_sub (hf : integrable f μ) (hg : integrable g μ) : + condexp_L1 hm μ (f - g) = condexp_L1 hm μ f - condexp_L1 hm μ g := +set_to_fun_sub _ hf hg + +lemma condexp_L1_of_ae_strongly_measurable' + (hfm : ae_strongly_measurable' m f μ) (hfi : integrable f μ) : + condexp_L1 hm μ f =ᵐ[μ] f := +begin + rw condexp_L1_eq hfi, + refine eventually_eq.trans _ (integrable.coe_fn_to_L1 hfi), + rw condexp_L1_clm_of_ae_strongly_measurable', + exact ae_strongly_measurable'.congr hfm (integrable.coe_fn_to_L1 hfi).symm, +end + +lemma condexp_L1_mono {E} [normed_lattice_add_comm_group E] [complete_space E] [normed_space ℝ E] + [ordered_smul ℝ E] {f g : α → E} + (hf : integrable f μ) (hg : integrable g μ) (hfg : f ≤ᵐ[μ] g) : + condexp_L1 hm μ f ≤ᵐ[μ] condexp_L1 hm μ g := +begin + rw coe_fn_le, + have h_nonneg : ∀ s, measurable_set s → μ s < ∞ → ∀ x : E, 0 ≤ x → 0 ≤ condexp_ind hm μ s x, + from λ s hs hμs x hx, condexp_ind_nonneg hs hμs.ne x hx, + exact set_to_fun_mono (dominated_fin_meas_additive_condexp_ind E hm μ) h_nonneg hf hg hfg, +end + +end condexp_L1 + +end measure_theory diff --git a/src/measure_theory/function/conditional_expectation/condexp_L2.lean b/src/measure_theory/function/conditional_expectation/condexp_L2.lean new file mode 100644 index 0000000000000..fc1d507a730ff --- /dev/null +++ b/src/measure_theory/function/conditional_expectation/condexp_L2.lean @@ -0,0 +1,518 @@ +/- +Copyright (c) 2021 Rémy Degenne. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Rémy Degenne +-/ +import analysis.inner_product_space.projection +import measure_theory.function.conditional_expectation.unique + +/-! # Conditional expectation in L2 + +This file contains one step of the construction of the conditional expectation, which is completed +in `measure_theory.function.conditional_expectation.basic`. See that file for a description of the +full process. + +We build the conditional expectation of an `L²` function, as an element of `L²`. This is the +orthogonal projection on the subspace of almost everywhere `m`-measurable functions. + +## Main definitions + +* `condexp_L2`: Conditional expectation of a function in L2 with respect to a sigma-algebra: it is +the orthogonal projection on the subspace `Lp_meas`. + +## Implementation notes + +Most of the results in this file are valid for a complete real normed space `F`. +However, some lemmas also use `𝕜 : is_R_or_C`: +* `condexp_L2` is defined only for an `inner_product_space` for now, and we use `𝕜` for its field. +* results about scalar multiplication are stated not only for `ℝ` but also for `𝕜` if we happen to + have `normed_space 𝕜 F`. + +-/ + +open topological_space filter continuous_linear_map +open_locale ennreal topology measure_theory + +namespace measure_theory + +variables {α E E' F G G' 𝕜 : Type*} {p : ℝ≥0∞} + [is_R_or_C 𝕜] -- 𝕜 for ℝ or ℂ + -- E for an inner product space + [normed_add_comm_group E] [inner_product_space 𝕜 E] [complete_space E] + -- E' for an inner product space on which we compute integrals + [normed_add_comm_group E'] [inner_product_space 𝕜 E'] + [complete_space E'] [normed_space ℝ E'] + -- F for a Lp submodule + [normed_add_comm_group F] [normed_space 𝕜 F] + -- G for a Lp add_subgroup + [normed_add_comm_group G] + -- G' for integrals on a Lp add_subgroup + [normed_add_comm_group G'] [normed_space ℝ G'] [complete_space G'] + +variables {m m0 : measurable_space α} {μ : measure α} {s t : set α} + +local notation `⟪`x`, `y`⟫` := @inner 𝕜 E _ x y +local notation `⟪`x`, `y`⟫₂` := @inner 𝕜 (α →₂[μ] E) _ x y + +variables (𝕜) +/-- Conditional expectation of a function in L2 with respect to a sigma-algebra -/ +noncomputable +def condexp_L2 (hm : m ≤ m0) : (α →₂[μ] E) →L[𝕜] (Lp_meas E 𝕜 m 2 μ) := +@orthogonal_projection 𝕜 (α →₂[μ] E) _ _ _ (Lp_meas E 𝕜 m 2 μ) + (by { haveI : fact (m ≤ m0) := ⟨hm⟩, exact infer_instance, }) +variables {𝕜} + +lemma ae_strongly_measurable'_condexp_L2 (hm : m ≤ m0) (f : α →₂[μ] E) : + ae_strongly_measurable' m (condexp_L2 𝕜 hm f) μ := +Lp_meas.ae_strongly_measurable' _ + +lemma integrable_on_condexp_L2_of_measure_ne_top (hm : m ≤ m0) (hμs : μ s ≠ ∞) (f : α →₂[μ] E) : + integrable_on (condexp_L2 𝕜 hm f) s μ := +integrable_on_Lp_of_measure_ne_top ((condexp_L2 𝕜 hm f) : α →₂[μ] E) + fact_one_le_two_ennreal.elim hμs + +lemma integrable_condexp_L2_of_is_finite_measure (hm : m ≤ m0) [is_finite_measure μ] + {f : α →₂[μ] E} : + integrable (condexp_L2 𝕜 hm f) μ := +integrable_on_univ.mp $ integrable_on_condexp_L2_of_measure_ne_top hm (measure_ne_top _ _) f + +lemma norm_condexp_L2_le_one (hm : m ≤ m0) : ‖@condexp_L2 α E 𝕜 _ _ _ _ _ _ μ hm‖ ≤ 1 := +by { haveI : fact (m ≤ m0) := ⟨hm⟩, exact orthogonal_projection_norm_le _, } + +lemma norm_condexp_L2_le (hm : m ≤ m0) (f : α →₂[μ] E) : ‖condexp_L2 𝕜 hm f‖ ≤ ‖f‖ := +((@condexp_L2 _ E 𝕜 _ _ _ _ _ _ μ hm).le_op_norm f).trans + (mul_le_of_le_one_left (norm_nonneg _) (norm_condexp_L2_le_one hm)) + +lemma snorm_condexp_L2_le (hm : m ≤ m0) (f : α →₂[μ] E) : + snorm (condexp_L2 𝕜 hm f) 2 μ ≤ snorm f 2 μ := +begin + rw [Lp_meas_coe, ← ennreal.to_real_le_to_real (Lp.snorm_ne_top _) (Lp.snorm_ne_top _), + ← Lp.norm_def, ← Lp.norm_def, submodule.norm_coe], + exact norm_condexp_L2_le hm f, +end + +lemma norm_condexp_L2_coe_le (hm : m ≤ m0) (f : α →₂[μ] E) : + ‖(condexp_L2 𝕜 hm f : α →₂[μ] E)‖ ≤ ‖f‖ := +begin + rw [Lp.norm_def, Lp.norm_def, ← Lp_meas_coe], + refine (ennreal.to_real_le_to_real _ (Lp.snorm_ne_top _)).mpr (snorm_condexp_L2_le hm f), + exact Lp.snorm_ne_top _, +end + +lemma inner_condexp_L2_left_eq_right (hm : m ≤ m0) {f g : α →₂[μ] E} : + ⟪(condexp_L2 𝕜 hm f : α →₂[μ] E), g⟫₂ = ⟪f, (condexp_L2 𝕜 hm g : α →₂[μ] E)⟫₂ := +by { haveI : fact (m ≤ m0) := ⟨hm⟩, exact inner_orthogonal_projection_left_eq_right _ f g, } + +lemma condexp_L2_indicator_of_measurable (hm : m ≤ m0) + (hs : measurable_set[m] s) (hμs : μ s ≠ ∞) (c : E) : + (condexp_L2 𝕜 hm (indicator_const_Lp 2 (hm s hs) hμs c) : α →₂[μ] E) + = indicator_const_Lp 2 (hm s hs) hμs c := +begin + rw condexp_L2, + haveI : fact (m ≤ m0) := ⟨hm⟩, + have h_mem : indicator_const_Lp 2 (hm s hs) hμs c ∈ Lp_meas E 𝕜 m 2 μ, + from mem_Lp_meas_indicator_const_Lp hm hs hμs, + let ind := (⟨indicator_const_Lp 2 (hm s hs) hμs c, h_mem⟩ : Lp_meas E 𝕜 m 2 μ), + have h_coe_ind : (ind : α →₂[μ] E) = indicator_const_Lp 2 (hm s hs) hμs c, by refl, + have h_orth_mem := orthogonal_projection_mem_subspace_eq_self ind, + rw [← h_coe_ind, h_orth_mem], +end + +lemma inner_condexp_L2_eq_inner_fun (hm : m ≤ m0) (f g : α →₂[μ] E) + (hg : ae_strongly_measurable' m g μ) : + ⟪(condexp_L2 𝕜 hm f : α →₂[μ] E), g⟫₂ = ⟪f, g⟫₂ := +begin + symmetry, + rw [← sub_eq_zero, ← inner_sub_left, condexp_L2], + simp only [mem_Lp_meas_iff_ae_strongly_measurable'.mpr hg, orthogonal_projection_inner_eq_zero], +end + +section real + +variables {hm : m ≤ m0} + +lemma integral_condexp_L2_eq_of_fin_meas_real (f : Lp 𝕜 2 μ) (hs : measurable_set[m] s) + (hμs : μ s ≠ ∞) : + ∫ x in s, condexp_L2 𝕜 hm f x ∂μ = ∫ x in s, f x ∂μ := +begin + rw ← L2.inner_indicator_const_Lp_one (hm s hs) hμs, + have h_eq_inner : ∫ x in s, condexp_L2 𝕜 hm f x ∂μ + = inner (indicator_const_Lp 2 (hm s hs) hμs (1 : 𝕜)) (condexp_L2 𝕜 hm f), + { rw L2.inner_indicator_const_Lp_one (hm s hs) hμs, + congr, }, + rw [h_eq_inner, ← inner_condexp_L2_left_eq_right, condexp_L2_indicator_of_measurable hm hs hμs], +end + +lemma lintegral_nnnorm_condexp_L2_le (hs : measurable_set[m] s) (hμs : μ s ≠ ∞) (f : Lp ℝ 2 μ) : + ∫⁻ x in s, ‖condexp_L2 ℝ hm f x‖₊ ∂μ ≤ ∫⁻ x in s, ‖f x‖₊ ∂μ := +begin + let h_meas := Lp_meas.ae_strongly_measurable' (condexp_L2 ℝ hm f), + let g := h_meas.some, + have hg_meas : strongly_measurable[m] g, from h_meas.some_spec.1, + have hg_eq : g =ᵐ[μ] condexp_L2 ℝ hm f, from h_meas.some_spec.2.symm, + have hg_eq_restrict : g =ᵐ[μ.restrict s] condexp_L2 ℝ hm f, from ae_restrict_of_ae hg_eq, + have hg_nnnorm_eq : (λ x, (‖g x‖₊ : ℝ≥0∞)) + =ᵐ[μ.restrict s] (λ x, (‖condexp_L2 ℝ hm f x‖₊ : ℝ≥0∞)), + { refine hg_eq_restrict.mono (λ x hx, _), + dsimp only, + rw hx, }, + rw lintegral_congr_ae hg_nnnorm_eq.symm, + refine lintegral_nnnorm_le_of_forall_fin_meas_integral_eq hm + (Lp.strongly_measurable f) _ _ _ _ hs hμs, + { exact integrable_on_Lp_of_measure_ne_top f fact_one_le_two_ennreal.elim hμs, }, + { exact hg_meas, }, + { rw [integrable_on, integrable_congr hg_eq_restrict], + exact integrable_on_condexp_L2_of_measure_ne_top hm hμs f, }, + { intros t ht hμt, + rw ← integral_condexp_L2_eq_of_fin_meas_real f ht hμt.ne, + exact set_integral_congr_ae (hm t ht) (hg_eq.mono (λ x hx _, hx)), }, +end + +lemma condexp_L2_ae_eq_zero_of_ae_eq_zero (hs : measurable_set[m] s) (hμs : μ s ≠ ∞) + {f : Lp ℝ 2 μ} (hf : f =ᵐ[μ.restrict s] 0) : + condexp_L2 ℝ hm f =ᵐ[μ.restrict s] 0 := +begin + suffices h_nnnorm_eq_zero : ∫⁻ x in s, ‖condexp_L2 ℝ hm f x‖₊ ∂μ = 0, + { rw lintegral_eq_zero_iff at h_nnnorm_eq_zero, + refine h_nnnorm_eq_zero.mono (λ x hx, _), + dsimp only at hx, + rw pi.zero_apply at hx ⊢, + { rwa [ennreal.coe_eq_zero, nnnorm_eq_zero] at hx, }, + { refine measurable.coe_nnreal_ennreal (measurable.nnnorm _), + rw Lp_meas_coe, + exact (Lp.strongly_measurable _).measurable }, }, + refine le_antisymm _ (zero_le _), + refine (lintegral_nnnorm_condexp_L2_le hs hμs f).trans (le_of_eq _), + rw lintegral_eq_zero_iff, + { refine hf.mono (λ x hx, _), + dsimp only, + rw hx, + simp, }, + { exact (Lp.strongly_measurable _).ennnorm, }, +end + +lemma lintegral_nnnorm_condexp_L2_indicator_le_real + (hs : measurable_set s) (hμs : μ s ≠ ∞) (ht : measurable_set[m] t) (hμt : μ t ≠ ∞) : + ∫⁻ a in t, ‖condexp_L2 ℝ hm (indicator_const_Lp 2 hs hμs (1 : ℝ)) a‖₊ ∂μ ≤ μ (s ∩ t) := +begin + refine (lintegral_nnnorm_condexp_L2_le ht hμt _).trans (le_of_eq _), + have h_eq : ∫⁻ x in t, ‖(indicator_const_Lp 2 hs hμs (1 : ℝ)) x‖₊ ∂μ + = ∫⁻ x in t, s.indicator (λ x, (1 : ℝ≥0∞)) x ∂μ, + { 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], + simp only [one_mul, set.univ_inter, measurable_set.univ, measure.restrict_apply], +end + +end real + +/-- `condexp_L2` commutes with taking inner products with constants. See the lemma +`condexp_L2_comp_continuous_linear_map` for a more general result about commuting with continuous +linear maps. -/ +lemma condexp_L2_const_inner (hm : m ≤ m0) (f : Lp E 2 μ) (c : E) : + condexp_L2 𝕜 hm (((Lp.mem_ℒp f).const_inner c).to_Lp (λ a, ⟪c, f a⟫)) + =ᵐ[μ] λ a, ⟪c, condexp_L2 𝕜 hm f a⟫ := +begin + rw Lp_meas_coe, + have h_mem_Lp : mem_ℒp (λ a, ⟪c, condexp_L2 𝕜 hm f a⟫) 2 μ, + { refine mem_ℒp.const_inner _ _, rw Lp_meas_coe, exact Lp.mem_ℒp _, }, + have h_eq : h_mem_Lp.to_Lp _ =ᵐ[μ] λ a, ⟪c, condexp_L2 𝕜 hm f a⟫, from h_mem_Lp.coe_fn_to_Lp, + refine eventually_eq.trans _ h_eq, + refine Lp.ae_eq_of_forall_set_integral_eq' 𝕜 hm _ _ two_ne_zero ennreal.coe_ne_top + (λ s hs hμs, integrable_on_condexp_L2_of_measure_ne_top hm hμs.ne _) _ _ _ _, + { intros s hs hμs, + rw [integrable_on, integrable_congr (ae_restrict_of_ae h_eq)], + exact (integrable_on_condexp_L2_of_measure_ne_top hm hμs.ne _).const_inner _, }, + { intros s hs hμs, + rw [← Lp_meas_coe, integral_condexp_L2_eq_of_fin_meas_real _ hs hμs.ne, + integral_congr_ae (ae_restrict_of_ae h_eq), Lp_meas_coe, + ← L2.inner_indicator_const_Lp_eq_set_integral_inner 𝕜 ↑(condexp_L2 𝕜 hm f) (hm s hs) c hμs.ne, + ← inner_condexp_L2_left_eq_right, condexp_L2_indicator_of_measurable, + L2.inner_indicator_const_Lp_eq_set_integral_inner 𝕜 f (hm s hs) c hμs.ne, + set_integral_congr_ae (hm s hs) + ((mem_ℒp.coe_fn_to_Lp ((Lp.mem_ℒp f).const_inner c)).mono (λ x hx hxs, hx))], }, + { rw ← Lp_meas_coe, exact Lp_meas.ae_strongly_measurable' _, }, + { refine ae_strongly_measurable'.congr _ h_eq.symm, + exact (Lp_meas.ae_strongly_measurable' _).const_inner _, }, +end + +/-- `condexp_L2` verifies the equality of integrals defining the conditional expectation. -/ +lemma integral_condexp_L2_eq (hm : m ≤ m0) + (f : Lp E' 2 μ) (hs : measurable_set[m] s) (hμs : μ s ≠ ∞) : + ∫ x in s, condexp_L2 𝕜 hm f x ∂μ = ∫ x in s, f x ∂μ := +begin + rw [← sub_eq_zero, Lp_meas_coe, ← integral_sub' + (integrable_on_Lp_of_measure_ne_top _ fact_one_le_two_ennreal.elim hμs) + (integrable_on_Lp_of_measure_ne_top _ fact_one_le_two_ennreal.elim hμs)], + refine integral_eq_zero_of_forall_integral_inner_eq_zero 𝕜 _ _ _, + { rw integrable_congr (ae_restrict_of_ae (Lp.coe_fn_sub ↑(condexp_L2 𝕜 hm f) f).symm), + exact integrable_on_Lp_of_measure_ne_top _ fact_one_le_two_ennreal.elim hμs, }, + intro c, + simp_rw [pi.sub_apply, inner_sub_right], + rw integral_sub + ((integrable_on_Lp_of_measure_ne_top _ fact_one_le_two_ennreal.elim hμs).const_inner c) + ((integrable_on_Lp_of_measure_ne_top _ fact_one_le_two_ennreal.elim hμs).const_inner c), + have h_ae_eq_f := mem_ℒp.coe_fn_to_Lp ((Lp.mem_ℒp f).const_inner c), + rw [← Lp_meas_coe, sub_eq_zero, + ← set_integral_congr_ae (hm s hs) ((condexp_L2_const_inner hm f c).mono (λ x hx _, hx)), + ← set_integral_congr_ae (hm s hs) (h_ae_eq_f.mono (λ x hx _, hx))], + exact integral_condexp_L2_eq_of_fin_meas_real _ hs hμs, +end + +variables {E'' 𝕜' : Type*} [is_R_or_C 𝕜'] [normed_add_comm_group E''] + [inner_product_space 𝕜' E''] [complete_space E''] [normed_space ℝ E''] + +variables (𝕜 𝕜') +lemma condexp_L2_comp_continuous_linear_map (hm : m ≤ m0) (T : E' →L[ℝ] E'') (f : α →₂[μ] E') : + (condexp_L2 𝕜' hm (T.comp_Lp f) : α →₂[μ] E'') =ᵐ[μ] T.comp_Lp (condexp_L2 𝕜 hm f : α →₂[μ] E') := +begin + refine Lp.ae_eq_of_forall_set_integral_eq' 𝕜' hm _ _ two_ne_zero ennreal.coe_ne_top + (λ s hs hμs, integrable_on_condexp_L2_of_measure_ne_top hm hμs.ne _) + (λ s hs hμs, integrable_on_Lp_of_measure_ne_top _ fact_one_le_two_ennreal.elim hμs.ne) + _ _ _, + { intros s hs hμs, + rw [T.set_integral_comp_Lp _ (hm s hs), + T.integral_comp_comm + (integrable_on_Lp_of_measure_ne_top _ fact_one_le_two_ennreal.elim hμs.ne), + ← Lp_meas_coe, ← Lp_meas_coe, integral_condexp_L2_eq hm f hs hμs.ne, + integral_condexp_L2_eq hm (T.comp_Lp f) hs hμs.ne, T.set_integral_comp_Lp _ (hm s hs), + T.integral_comp_comm + (integrable_on_Lp_of_measure_ne_top f fact_one_le_two_ennreal.elim hμs.ne)], }, + { rw ← Lp_meas_coe, exact Lp_meas.ae_strongly_measurable' _, }, + { have h_coe := T.coe_fn_comp_Lp (condexp_L2 𝕜 hm f : α →₂[μ] E'), + rw ← eventually_eq at h_coe, + refine ae_strongly_measurable'.congr _ h_coe.symm, + exact (Lp_meas.ae_strongly_measurable' (condexp_L2 𝕜 hm f)).continuous_comp T.continuous, }, +end +variables {𝕜 𝕜'} + +section condexp_L2_indicator + +variables (𝕜) +lemma condexp_L2_indicator_ae_eq_smul (hm : m ≤ m0) (hs : measurable_set s) (hμs : μ s ≠ ∞) + (x : E') : + condexp_L2 𝕜 hm (indicator_const_Lp 2 hs hμs x) + =ᵐ[μ] λ a, (condexp_L2 ℝ hm (indicator_const_Lp 2 hs hμs (1 : ℝ)) a) • x := +begin + rw indicator_const_Lp_eq_to_span_singleton_comp_Lp hs hμs x, + have h_comp := condexp_L2_comp_continuous_linear_map ℝ 𝕜 hm (to_span_singleton ℝ x) + (indicator_const_Lp 2 hs hμs (1 : ℝ)), + rw ← Lp_meas_coe at h_comp, + refine h_comp.trans _, + exact (to_span_singleton ℝ x).coe_fn_comp_Lp _, +end + +lemma condexp_L2_indicator_eq_to_span_singleton_comp (hm : m ≤ m0) (hs : measurable_set s) + (hμs : μ s ≠ ∞) (x : E') : + (condexp_L2 𝕜 hm (indicator_const_Lp 2 hs hμs x) : α →₂[μ] E') + = (to_span_singleton ℝ x).comp_Lp (condexp_L2 ℝ hm (indicator_const_Lp 2 hs hμs (1 : ℝ))) := +begin + ext1, + rw ← Lp_meas_coe, + refine (condexp_L2_indicator_ae_eq_smul 𝕜 hm hs hμs x).trans _, + have h_comp := (to_span_singleton ℝ x).coe_fn_comp_Lp + (condexp_L2 ℝ hm (indicator_const_Lp 2 hs hμs (1 : ℝ)) : α →₂[μ] ℝ), + rw ← eventually_eq at h_comp, + refine eventually_eq.trans _ h_comp.symm, + refine eventually_of_forall (λ y, _), + refl, +end + +variables {𝕜} + +lemma set_lintegral_nnnorm_condexp_L2_indicator_le (hm : m ≤ m0) (hs : measurable_set s) + (hμs : μ s ≠ ∞) (x : E') {t : set α} (ht : measurable_set[m] t) (hμt : μ t ≠ ∞) : + ∫⁻ a in t, ‖condexp_L2 𝕜 hm (indicator_const_Lp 2 hs hμs x) a‖₊ ∂μ ≤ μ (s ∩ t) * ‖x‖₊ := +calc ∫⁻ a in t, ‖condexp_L2 𝕜 hm (indicator_const_Lp 2 hs hμs x) a‖₊ ∂μ + = ∫⁻ a in t, ‖(condexp_L2 ℝ hm (indicator_const_Lp 2 hs hμs (1 : ℝ)) a) • x‖₊ ∂μ : +set_lintegral_congr_fun (hm t ht) + ((condexp_L2_indicator_ae_eq_smul 𝕜 hm hs hμs x).mono (λ a ha hat, by rw ha)) +... = ∫⁻ a in t, ‖condexp_L2 ℝ hm (indicator_const_Lp 2 hs hμs (1 : ℝ)) a‖₊ ∂μ * ‖x‖₊ : +begin + simp_rw [nnnorm_smul, ennreal.coe_mul], + rw [lintegral_mul_const, Lp_meas_coe], + exact (Lp.strongly_measurable _).ennnorm +end +... ≤ μ (s ∩ t) * ‖x‖₊ : + mul_le_mul_right' (lintegral_nnnorm_condexp_L2_indicator_le_real hs hμs ht hμt) _ + +lemma lintegral_nnnorm_condexp_L2_indicator_le (hm : m ≤ m0) (hs : measurable_set s) + (hμs : μ s ≠ ∞) (x : E') [sigma_finite (μ.trim hm)] : + ∫⁻ a, ‖condexp_L2 𝕜 hm (indicator_const_Lp 2 hs hμs x) a‖₊ ∂μ ≤ μ s * ‖x‖₊ := +begin + refine lintegral_le_of_forall_fin_meas_le' hm (μ s * ‖x‖₊) _ (λ t ht hμt, _), + { rw Lp_meas_coe, + exact (Lp.ae_strongly_measurable _).ennnorm }, + refine (set_lintegral_nnnorm_condexp_L2_indicator_le hm hs hμs x ht hμt).trans _, + exact mul_le_mul_right' (measure_mono (set.inter_subset_left _ _)) _ +end + +/-- If the measure `μ.trim hm` is sigma-finite, then the conditional expectation of a measurable set +with finite measure is integrable. -/ +lemma integrable_condexp_L2_indicator (hm : m ≤ m0) [sigma_finite (μ.trim hm)] + (hs : measurable_set s) (hμs : μ s ≠ ∞) (x : E') : + integrable (condexp_L2 𝕜 hm (indicator_const_Lp 2 hs hμs x)) μ := +begin + refine integrable_of_forall_fin_meas_le' hm (μ s * ‖x‖₊) + (ennreal.mul_lt_top hμs ennreal.coe_ne_top) _ _, + { rw Lp_meas_coe, exact Lp.ae_strongly_measurable _, }, + { refine λ t ht hμt, (set_lintegral_nnnorm_condexp_L2_indicator_le hm hs hμs x ht hμt).trans _, + exact mul_le_mul_right' (measure_mono (set.inter_subset_left _ _)) _, }, +end + +end condexp_L2_indicator + +section condexp_ind_smul + +variables [normed_space ℝ G] {hm : m ≤ m0} + +/-- Conditional expectation of the indicator of a measurable set with finite measure, in L2. -/ +noncomputable +def condexp_ind_smul (hm : m ≤ m0) (hs : measurable_set s) (hμs : μ s ≠ ∞) (x : G) : Lp G 2 μ := +(to_span_singleton ℝ x).comp_LpL 2 μ (condexp_L2 ℝ hm (indicator_const_Lp 2 hs hμs (1 : ℝ))) + +lemma ae_strongly_measurable'_condexp_ind_smul + (hm : m ≤ m0) (hs : measurable_set s) (hμs : μ s ≠ ∞) (x : G) : + ae_strongly_measurable' m (condexp_ind_smul hm hs hμs x) μ := +begin + have h : ae_strongly_measurable' m (condexp_L2 ℝ hm (indicator_const_Lp 2 hs hμs (1 : ℝ))) μ, + from ae_strongly_measurable'_condexp_L2 _ _, + rw condexp_ind_smul, + suffices : ae_strongly_measurable' m + ((to_span_singleton ℝ x) ∘ (condexp_L2 ℝ hm (indicator_const_Lp 2 hs hμs (1 : ℝ)))) μ, + { refine ae_strongly_measurable'.congr this _, + refine eventually_eq.trans _ (coe_fn_comp_LpL _ _).symm, + rw Lp_meas_coe, }, + exact ae_strongly_measurable'.continuous_comp (to_span_singleton ℝ x).continuous h, +end + +lemma condexp_ind_smul_add (hs : measurable_set s) (hμs : μ s ≠ ∞) (x y : G) : + condexp_ind_smul hm hs hμs (x + y) + = condexp_ind_smul hm hs hμs x + condexp_ind_smul hm hs hμs y := +by { simp_rw [condexp_ind_smul], rw [to_span_singleton_add, add_comp_LpL, add_apply], } + +lemma condexp_ind_smul_smul (hs : measurable_set s) (hμs : μ s ≠ ∞) (c : ℝ) (x : G) : + condexp_ind_smul hm hs hμs (c • x) = c • condexp_ind_smul hm hs hμs x := +by { simp_rw [condexp_ind_smul], rw [to_span_singleton_smul, smul_comp_LpL, smul_apply], } + +lemma condexp_ind_smul_smul' [normed_space ℝ F] [smul_comm_class ℝ 𝕜 F] (hs : measurable_set s) + (hμs : μ s ≠ ∞) (c : 𝕜) (x : F) : + condexp_ind_smul hm hs hμs (c • x) = c • condexp_ind_smul hm hs hμs x := +by rw [condexp_ind_smul, condexp_ind_smul, to_span_singleton_smul', + (to_span_singleton ℝ x).smul_comp_LpL c, smul_apply] + +lemma condexp_ind_smul_ae_eq_smul (hm : m ≤ m0) (hs : measurable_set s) (hμs : μ s ≠ ∞) (x : G) : + condexp_ind_smul hm hs hμs x + =ᵐ[μ] λ a, (condexp_L2 ℝ hm (indicator_const_Lp 2 hs hμs (1 : ℝ)) a) • x := +(to_span_singleton ℝ x).coe_fn_comp_LpL _ + +lemma set_lintegral_nnnorm_condexp_ind_smul_le (hm : m ≤ m0) (hs : measurable_set s) + (hμs : μ s ≠ ∞) (x : G) {t : set α} (ht : measurable_set[m] t) (hμt : μ t ≠ ∞) : + ∫⁻ a in t, ‖condexp_ind_smul hm hs hμs x a‖₊ ∂μ ≤ μ (s ∩ t) * ‖x‖₊ := +calc ∫⁻ a in t, ‖condexp_ind_smul hm hs hμs x a‖₊ ∂μ + = ∫⁻ a in t, ‖condexp_L2 ℝ hm (indicator_const_Lp 2 hs hμs (1 : ℝ)) a • x‖₊ ∂μ : +set_lintegral_congr_fun (hm t ht) + ((condexp_ind_smul_ae_eq_smul hm hs hμs x).mono (λ a ha hat, by rw ha )) +... = ∫⁻ a in t, ‖condexp_L2 ℝ hm (indicator_const_Lp 2 hs hμs (1 : ℝ)) a‖₊ ∂μ * ‖x‖₊ : +begin + simp_rw [nnnorm_smul, ennreal.coe_mul], + rw [lintegral_mul_const, Lp_meas_coe], + exact (Lp.strongly_measurable _).ennnorm +end +... ≤ μ (s ∩ t) * ‖x‖₊ : + mul_le_mul_right' (lintegral_nnnorm_condexp_L2_indicator_le_real hs hμs ht hμt) _ + +lemma lintegral_nnnorm_condexp_ind_smul_le (hm : m ≤ m0) (hs : measurable_set s) + (hμs : μ s ≠ ∞) (x : G) [sigma_finite (μ.trim hm)] : + ∫⁻ a, ‖condexp_ind_smul hm hs hμs x a‖₊ ∂μ ≤ μ s * ‖x‖₊ := +begin + refine lintegral_le_of_forall_fin_meas_le' hm (μ s * ‖x‖₊) _ (λ t ht hμt, _), + { exact (Lp.ae_strongly_measurable _).ennnorm }, + refine (set_lintegral_nnnorm_condexp_ind_smul_le hm hs hμs x ht hμt).trans _, + exact mul_le_mul_right' (measure_mono (set.inter_subset_left _ _)) _ +end + +/-- If the measure `μ.trim hm` is sigma-finite, then the conditional expectation of a measurable set +with finite measure is integrable. -/ +lemma integrable_condexp_ind_smul (hm : m ≤ m0) [sigma_finite (μ.trim hm)] + (hs : measurable_set s) (hμs : μ s ≠ ∞) (x : G) : + integrable (condexp_ind_smul hm hs hμs x) μ := +begin + refine integrable_of_forall_fin_meas_le' hm (μ s * ‖x‖₊) + (ennreal.mul_lt_top hμs ennreal.coe_ne_top) _ _, + { exact Lp.ae_strongly_measurable _, }, + { refine λ t ht hμt, (set_lintegral_nnnorm_condexp_ind_smul_le hm hs hμs x ht hμt).trans _, + exact mul_le_mul_right' (measure_mono (set.inter_subset_left _ _)) _, }, +end + +lemma condexp_ind_smul_empty {x : G} : + condexp_ind_smul hm measurable_set.empty + ((@measure_empty _ _ μ).le.trans_lt ennreal.coe_lt_top).ne x = 0 := +begin + rw [condexp_ind_smul, indicator_const_empty], + simp only [coe_fn_coe_base, submodule.coe_zero, continuous_linear_map.map_zero], +end + +lemma set_integral_condexp_L2_indicator (hs : measurable_set[m] s) (ht : measurable_set t) + (hμs : μ s ≠ ∞) (hμt : μ t ≠ ∞) : + ∫ x in s, (condexp_L2 ℝ hm (indicator_const_Lp 2 ht hμt (1 : ℝ))) x ∂μ = (μ (t ∩ s)).to_real := +calc ∫ x in s, (condexp_L2 ℝ hm (indicator_const_Lp 2 ht hμt (1 : ℝ))) x ∂μ + = ∫ x in s, indicator_const_Lp 2 ht hμt (1 : ℝ) x ∂μ : + @integral_condexp_L2_eq + α _ ℝ _ _ _ _ _ _ _ _ _ hm (indicator_const_Lp 2 ht hμt (1 : ℝ)) hs hμs +... = (μ (t ∩ s)).to_real • 1 : set_integral_indicator_const_Lp (hm s hs) ht hμt (1 : ℝ) +... = (μ (t ∩ s)).to_real : by rw [smul_eq_mul, mul_one] + +lemma set_integral_condexp_ind_smul (hs : measurable_set[m] s) (ht : measurable_set t) + (hμs : μ s ≠ ∞) (hμt : μ t ≠ ∞) (x : G') : + ∫ a in s, (condexp_ind_smul hm ht hμt x) a ∂μ = (μ (t ∩ s)).to_real • x := +calc ∫ a in s, (condexp_ind_smul hm ht hμt x) a ∂μ + = (∫ a in s, (condexp_L2 ℝ hm (indicator_const_Lp 2 ht hμt (1 : ℝ)) a • x) ∂μ) : + set_integral_congr_ae (hm s hs) ((condexp_ind_smul_ae_eq_smul hm ht hμt x).mono (λ x hx hxs, hx)) +... = (∫ a in s, condexp_L2 ℝ hm (indicator_const_Lp 2 ht hμt (1 : ℝ)) a ∂μ) • x : + integral_smul_const _ x +... = (μ (t ∩ s)).to_real • x : + by rw set_integral_condexp_L2_indicator hs ht hμs hμt + +lemma condexp_L2_indicator_nonneg (hm : m ≤ m0) (hs : measurable_set s) (hμs : μ s ≠ ∞) + [sigma_finite (μ.trim hm)] : + 0 ≤ᵐ[μ] condexp_L2 ℝ hm (indicator_const_Lp 2 hs hμs (1 : ℝ)) := +begin + have h : ae_strongly_measurable' m (condexp_L2 ℝ hm (indicator_const_Lp 2 hs hμs (1 : ℝ))) μ, + from ae_strongly_measurable'_condexp_L2 _ _, + refine eventually_le.trans_eq _ h.ae_eq_mk.symm, + refine @ae_le_of_ae_le_trim _ _ _ _ _ _ hm _ _ _, + refine ae_nonneg_of_forall_set_integral_nonneg_of_sigma_finite _ _, + { intros t ht hμt, + refine @integrable.integrable_on _ _ m _ _ _ _ _, + refine integrable.trim hm _ _, + { rw integrable_congr h.ae_eq_mk.symm, + exact integrable_condexp_L2_indicator hm hs hμs _, }, + { exact h.strongly_measurable_mk, }, }, + { intros t ht hμt, + rw ← set_integral_trim hm h.strongly_measurable_mk ht, + have h_ae : ∀ᵐ x ∂μ, x ∈ t → h.mk _ x = condexp_L2 ℝ hm (indicator_const_Lp 2 hs hμs (1 : ℝ)) x, + { filter_upwards [h.ae_eq_mk] with x hx, + exact λ _, hx.symm, }, + rw [set_integral_congr_ae (hm t ht) h_ae, + set_integral_condexp_L2_indicator ht hs ((le_trim hm).trans_lt hμt).ne hμs], + exact ennreal.to_real_nonneg, }, +end + +lemma condexp_ind_smul_nonneg {E} [normed_lattice_add_comm_group E] [normed_space ℝ E] + [ordered_smul ℝ E] [sigma_finite (μ.trim hm)] + (hs : measurable_set s) (hμs : μ s ≠ ∞) (x : E) (hx : 0 ≤ x) : + 0 ≤ᵐ[μ] condexp_ind_smul hm hs hμs x := +begin + refine eventually_le.trans_eq _ (condexp_ind_smul_ae_eq_smul hm hs hμs x).symm, + filter_upwards [condexp_L2_indicator_nonneg hm hs hμs] with a ha, + exact smul_nonneg ha hx, +end + +end condexp_ind_smul + +end measure_theory diff --git a/src/measure_theory/function/conditional_expectation/unique.lean b/src/measure_theory/function/conditional_expectation/unique.lean new file mode 100644 index 0000000000000..7d4488dcb41de --- /dev/null +++ b/src/measure_theory/function/conditional_expectation/unique.lean @@ -0,0 +1,213 @@ +/- +Copyright (c) 2021 Rémy Degenne. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Rémy Degenne +-/ +import measure_theory.function.ae_eq_of_integral +import measure_theory.function.conditional_expectation.ae_measurable + +/-! +# Uniqueness of the conditional expectation + +Two Lp functions `f, g` which are almost everywhere strongly measurable with respect to a σ-algebra +`m` and verify `∫ x in s, f x ∂μ = ∫ x in s, g x ∂μ` for all `m`-measurable sets `s` are equal +almost everywhere. This proves the uniqueness of the conditional expectation, which is not yet +defined in this file but is introduced in `measure_theory.function.conditional_expectation.basic`. + +## Main statements + +* `Lp.ae_eq_of_forall_set_integral_eq'`: two `Lp` functions verifying the equality of integrals + defining the conditional expectation are equal. +* `ae_eq_of_forall_set_integral_eq_of_sigma_finite'`: two functions verifying the equality of + integrals defining the conditional expectation are equal almost everywhere. + Requires `[sigma_finite (μ.trim hm)]`. + +-/ + +open_locale ennreal measure_theory + +namespace measure_theory + +variables {α E' F' 𝕜 : Type*} {p : ℝ≥0∞} + {m m0 : measurable_space α} {μ : measure α} + [is_R_or_C 𝕜] -- 𝕜 for ℝ or ℂ + -- E' for an inner product space on which we compute integrals + [normed_add_comm_group E'] [inner_product_space 𝕜 E'] + [complete_space E'] [normed_space ℝ E'] + -- F' for integrals on a Lp submodule + [normed_add_comm_group F'] [normed_space 𝕜 F'] [normed_space ℝ F'] [complete_space F'] + +section uniqueness_of_conditional_expectation + +/-! ## Uniqueness of the conditional expectation -/ + +lemma Lp_meas.ae_eq_zero_of_forall_set_integral_eq_zero + (hm : m ≤ m0) (f : Lp_meas E' 𝕜 m p μ) (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) + (hf_int_finite : ∀ s, measurable_set[m] s → μ s < ∞ → integrable_on f s μ) + (hf_zero : ∀ s : set α, measurable_set[m] s → μ s < ∞ → ∫ x in s, f x ∂μ = 0) : + f =ᵐ[μ] 0 := +begin + obtain ⟨g, hg_sm, hfg⟩ := Lp_meas.ae_fin_strongly_measurable' hm f hp_ne_zero hp_ne_top, + refine hfg.trans _, + refine ae_eq_zero_of_forall_set_integral_eq_of_fin_strongly_measurable_trim hm _ _ hg_sm, + { intros s hs hμs, + have hfg_restrict : f =ᵐ[μ.restrict s] g, from ae_restrict_of_ae hfg, + rw [integrable_on, integrable_congr hfg_restrict.symm], + exact hf_int_finite s hs hμs, }, + { intros s hs hμs, + have hfg_restrict : f =ᵐ[μ.restrict s] g, from ae_restrict_of_ae hfg, + rw integral_congr_ae hfg_restrict.symm, + exact hf_zero s hs hμs, }, +end + +include 𝕜 +variables (𝕜) + +lemma Lp.ae_eq_zero_of_forall_set_integral_eq_zero' + (hm : m ≤ m0) (f : Lp E' p μ) (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) + (hf_int_finite : ∀ s, measurable_set[m] s → μ s < ∞ → integrable_on f s μ) + (hf_zero : ∀ s : set α, measurable_set[m] s → μ s < ∞ → ∫ x in s, f x ∂μ = 0) + (hf_meas : ae_strongly_measurable' m f μ) : + f =ᵐ[μ] 0 := +begin + let f_meas : Lp_meas E' 𝕜 m p μ := ⟨f, hf_meas⟩, + have hf_f_meas : f =ᵐ[μ] f_meas, by simp only [coe_fn_coe_base', subtype.coe_mk], + refine hf_f_meas.trans _, + refine Lp_meas.ae_eq_zero_of_forall_set_integral_eq_zero hm f_meas hp_ne_zero hp_ne_top _ _, + { intros s hs hμs, + have hfg_restrict : f =ᵐ[μ.restrict s] f_meas, from ae_restrict_of_ae hf_f_meas, + rw [integrable_on, integrable_congr hfg_restrict.symm], + exact hf_int_finite s hs hμs, }, + { intros s hs hμs, + have hfg_restrict : f =ᵐ[μ.restrict s] f_meas, from ae_restrict_of_ae hf_f_meas, + rw integral_congr_ae hfg_restrict.symm, + exact hf_zero s hs hμs, }, +end + +/-- **Uniqueness of the conditional expectation** -/ +lemma Lp.ae_eq_of_forall_set_integral_eq' + (hm : m ≤ m0) (f g : Lp E' p μ) (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) + (hf_int_finite : ∀ s, measurable_set[m] s → μ s < ∞ → integrable_on f s μ) + (hg_int_finite : ∀ s, measurable_set[m] s → μ s < ∞ → integrable_on g s μ) + (hfg : ∀ s : set α, measurable_set[m] s → μ s < ∞ → ∫ x in s, f x ∂μ = ∫ x in s, g x ∂μ) + (hf_meas : ae_strongly_measurable' m f μ) (hg_meas : ae_strongly_measurable' m g μ) : + f =ᵐ[μ] g := +begin + suffices h_sub : ⇑(f-g) =ᵐ[μ] 0, + by { rw ← sub_ae_eq_zero, exact (Lp.coe_fn_sub f g).symm.trans h_sub, }, + have hfg' : ∀ s : set α, measurable_set[m] s → μ s < ∞ → ∫ x in s, (f - g) x ∂μ = 0, + { intros s hs hμs, + rw integral_congr_ae (ae_restrict_of_ae (Lp.coe_fn_sub f g)), + rw integral_sub' (hf_int_finite s hs hμs) (hg_int_finite s hs hμs), + exact sub_eq_zero.mpr (hfg s hs hμs), }, + have hfg_int : ∀ s, measurable_set[m] s → μ s < ∞ → integrable_on ⇑(f-g) s μ, + { intros s hs hμs, + rw [integrable_on, integrable_congr (ae_restrict_of_ae (Lp.coe_fn_sub f g))], + exact (hf_int_finite s hs hμs).sub (hg_int_finite s hs hμs), }, + have hfg_meas : ae_strongly_measurable' m ⇑(f - g) μ, + from ae_strongly_measurable'.congr (hf_meas.sub hg_meas) (Lp.coe_fn_sub f g).symm, + exact Lp.ae_eq_zero_of_forall_set_integral_eq_zero' 𝕜 hm (f-g) hp_ne_zero hp_ne_top hfg_int hfg' + hfg_meas, +end + +variables {𝕜} +omit 𝕜 + +lemma ae_eq_of_forall_set_integral_eq_of_sigma_finite' (hm : m ≤ m0) [sigma_finite (μ.trim hm)] + {f g : α → F'} + (hf_int_finite : ∀ s, measurable_set[m] s → μ s < ∞ → integrable_on f s μ) + (hg_int_finite : ∀ s, measurable_set[m] s → μ s < ∞ → integrable_on g s μ) + (hfg_eq : ∀ s : set α, measurable_set[m] s → μ s < ∞ → ∫ x in s, f x ∂μ = ∫ x in s, g x ∂μ) + (hfm : ae_strongly_measurable' m f μ) (hgm : ae_strongly_measurable' m g μ) : + f =ᵐ[μ] g := +begin + rw ← ae_eq_trim_iff_of_ae_strongly_measurable' hm hfm hgm, + have hf_mk_int_finite : ∀ s, measurable_set[m] s → μ.trim hm s < ∞ → + @integrable_on _ _ m _ (hfm.mk f) s (μ.trim hm), + { intros s hs hμs, + rw trim_measurable_set_eq hm hs at hμs, + rw [integrable_on, restrict_trim hm _ hs], + refine integrable.trim hm _ hfm.strongly_measurable_mk, + exact integrable.congr (hf_int_finite s hs hμs) (ae_restrict_of_ae hfm.ae_eq_mk), }, + have hg_mk_int_finite : ∀ s, measurable_set[m] s → μ.trim hm s < ∞ → + @integrable_on _ _ m _ (hgm.mk g) s (μ.trim hm), + { intros s hs hμs, + rw trim_measurable_set_eq hm hs at hμs, + rw [integrable_on, restrict_trim hm _ hs], + refine integrable.trim hm _ hgm.strongly_measurable_mk, + exact integrable.congr (hg_int_finite s hs hμs) (ae_restrict_of_ae hgm.ae_eq_mk), }, + have hfg_mk_eq : ∀ s : set α, measurable_set[m] s → μ.trim hm s < ∞ → + ∫ x in s, (hfm.mk f x) ∂(μ.trim hm) = ∫ x in s, (hgm.mk g x) ∂(μ.trim hm), + { intros s hs hμs, + rw trim_measurable_set_eq hm hs at hμs, + rw [restrict_trim hm _ hs, ← integral_trim hm hfm.strongly_measurable_mk, + ← integral_trim hm hgm.strongly_measurable_mk, + integral_congr_ae (ae_restrict_of_ae hfm.ae_eq_mk.symm), + integral_congr_ae (ae_restrict_of_ae hgm.ae_eq_mk.symm)], + exact hfg_eq s hs hμs, }, + exact ae_eq_of_forall_set_integral_eq_of_sigma_finite hf_mk_int_finite hg_mk_int_finite hfg_mk_eq, +end + +end uniqueness_of_conditional_expectation + + +section integral_norm_le + +variables {s : set α} + +/-- Let `m` be a sub-σ-algebra of `m0`, `f` a `m0`-measurable function and `g` a `m`-measurable +function, such that their integrals coincide on `m`-measurable sets with finite measure. +Then `∫ x in s, ‖g x‖ ∂μ ≤ ∫ x in s, ‖f x‖ ∂μ` on all `m`-measurable sets with finite measure. -/ +lemma integral_norm_le_of_forall_fin_meas_integral_eq (hm : m ≤ m0) {f g : α → ℝ} + (hf : strongly_measurable f) (hfi : integrable_on f s μ) + (hg : strongly_measurable[m] g) (hgi : integrable_on g s μ) + (hgf : ∀ t, measurable_set[m] t → μ t < ∞ → ∫ x in t, g x ∂μ = ∫ x in t, f x ∂μ) + (hs : measurable_set[m] s) (hμs : μ s ≠ ∞) : + ∫ x in s, ‖g x‖ ∂μ ≤ ∫ x in s, ‖f x‖ ∂μ := +begin + rw [integral_norm_eq_pos_sub_neg hgi, integral_norm_eq_pos_sub_neg hfi], + have h_meas_nonneg_g : measurable_set[m] {x | 0 ≤ g x}, + from (@strongly_measurable_const _ _ m _ _).measurable_set_le hg, + have h_meas_nonneg_f : measurable_set {x | 0 ≤ f x}, + from strongly_measurable_const.measurable_set_le hf, + have h_meas_nonpos_g : measurable_set[m] {x | g x ≤ 0}, + from hg.measurable_set_le (@strongly_measurable_const _ _ m _ _), + have h_meas_nonpos_f : measurable_set {x | f x ≤ 0}, + from hf.measurable_set_le strongly_measurable_const, + refine sub_le_sub _ _, + { rw [measure.restrict_restrict (hm _ h_meas_nonneg_g), + measure.restrict_restrict h_meas_nonneg_f, + hgf _ (@measurable_set.inter α m _ _ h_meas_nonneg_g hs) + ((measure_mono (set.inter_subset_right _ _)).trans_lt (lt_top_iff_ne_top.mpr hμs)), + ← measure.restrict_restrict (hm _ h_meas_nonneg_g), + ← measure.restrict_restrict h_meas_nonneg_f], + exact set_integral_le_nonneg (hm _ h_meas_nonneg_g) hf hfi, }, + { rw [measure.restrict_restrict (hm _ h_meas_nonpos_g), + measure.restrict_restrict h_meas_nonpos_f, + hgf _ (@measurable_set.inter α m _ _ h_meas_nonpos_g hs) + ((measure_mono (set.inter_subset_right _ _)).trans_lt (lt_top_iff_ne_top.mpr hμs)), + ← measure.restrict_restrict (hm _ h_meas_nonpos_g), + ← measure.restrict_restrict h_meas_nonpos_f], + exact set_integral_nonpos_le (hm _ h_meas_nonpos_g) hf hfi, }, +end + +/-- Let `m` be a sub-σ-algebra of `m0`, `f` a `m0`-measurable function and `g` a `m`-measurable +function, such that their integrals coincide on `m`-measurable sets with finite measure. +Then `∫⁻ x in s, ‖g x‖₊ ∂μ ≤ ∫⁻ x in s, ‖f x‖₊ ∂μ` on all `m`-measurable sets with finite +measure. -/ +lemma lintegral_nnnorm_le_of_forall_fin_meas_integral_eq (hm : m ≤ m0) {f g : α → ℝ} + (hf : strongly_measurable f) (hfi : integrable_on f s μ) + (hg : strongly_measurable[m] g) (hgi : integrable_on g s μ) + (hgf : ∀ t, measurable_set[m] t → μ t < ∞ → ∫ x in t, g x ∂μ = ∫ x in t, f x ∂μ) + (hs : measurable_set[m] s) (hμs : μ s ≠ ∞) : + ∫⁻ x in s, ‖g x‖₊ ∂μ ≤ ∫⁻ x in s, ‖f x‖₊ ∂μ := +begin + rw [← of_real_integral_norm_eq_lintegral_nnnorm hfi, + ← of_real_integral_norm_eq_lintegral_nnnorm hgi, ennreal.of_real_le_of_real_iff], + { exact integral_norm_le_of_forall_fin_meas_integral_eq hm hf hfi hg hgi hgf hs hμs, }, + { exact integral_nonneg (λ x, norm_nonneg _), }, +end + +end integral_norm_le + +end measure_theory From 9fb8964792b4237dac6200193a0d533f1b3f7423 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 12 Jun 2023 19:52:45 +0000 Subject: [PATCH 1175/1291] refactor(field_theory/splitting_field/is_splitting_field): use `root_set` (#19179) Co-authored-by: Eric Rodriguez --- src/field_theory/finite/galois_field.lean | 4 +- src/field_theory/galois.lean | 9 ++-- src/field_theory/normal.lean | 7 ++- src/field_theory/polynomial_galois_group.lean | 2 +- .../splitting_field/construction.lean | 38 +++++++-------- .../splitting_field/is_splitting_field.lean | 28 ++++++----- src/number_theory/cyclotomic/basic.lean | 47 ++++++++++--------- 7 files changed, 66 insertions(+), 69 deletions(-) diff --git a/src/field_theory/finite/galois_field.lean b/src/field_theory/finite/galois_field.lean index bb674db9bbb99..967ee213f8185 100644 --- a/src/field_theory/finite/galois_field.lean +++ b/src/field_theory/finite/galois_field.lean @@ -43,11 +43,11 @@ instance finite_field.has_sub.sub.polynomial.is_splitting_field (K F : Type*) [f rw [←splits_id_iff_splits, splits_iff_card_roots, polynomial.map_sub, polynomial.map_pow, map_X, h, finite_field.roots_X_pow_card_sub_X K, ←finset.card_def, finset.card_univ], end, - adjoin_roots := + adjoin_root_set := begin classical, transitivity algebra.adjoin F ((roots (X ^ (fintype.card K) - X : K[X])).to_finset : set K), - { simp only [polynomial.map_pow, map_X, polynomial.map_sub], }, + { simp only [root_set, polynomial.map_pow, map_X, polynomial.map_sub], }, { rw [finite_field.roots_X_pow_card_sub_X, val_to_finset, coe_univ, algebra.adjoin_univ], } end } diff --git a/src/field_theory/galois.lean b/src/field_theory/galois.lean index 5b4cc08749054..1efa971954436 100644 --- a/src/field_theory/galois.lean +++ b/src/field_theory/galois.lean @@ -297,11 +297,8 @@ begin rw [eq_top_iff, ←intermediate_field.top_to_subalgebra, ←h1], rw intermediate_field.adjoin_simple_to_subalgebra_of_integral (integral F α), apply algebra.adjoin_mono, - rw [set.singleton_subset_iff, finset.mem_coe, multiset.mem_to_finset, polynomial.mem_roots], - { dsimp only [polynomial.is_root], - rw [polynomial.eval_map, ←polynomial.aeval_def], - exact minpoly.aeval _ _ }, - { exact polynomial.map_ne_zero (minpoly.ne_zero (integral F α)) } + rw [set.singleton_subset_iff, polynomial.mem_root_set], + exact ⟨minpoly.ne_zero (integral F α), minpoly.aeval _ _⟩ end lemma of_fixed_field_eq_bot [finite_dimensional F E] @@ -371,7 +368,7 @@ begin let s := (p.map (algebra_map F E)).roots.to_finset, have adjoin_root : intermediate_field.adjoin F ↑s = ⊤, { apply intermediate_field.to_subalgebra_injective, - rw [intermediate_field.top_to_subalgebra, ←top_le_iff, ←sp.adjoin_roots], + rw [intermediate_field.top_to_subalgebra, ←top_le_iff, ←sp.adjoin_root_set], apply intermediate_field.algebra_adjoin_le_adjoin, }, let P : intermediate_field F E → Prop := λ K, fintype.card (K →ₐ[F] E) = finrank F K, suffices : P (intermediate_field.adjoin F ↑s), diff --git a/src/field_theory/normal.lean b/src/field_theory/normal.lean index 6b43eea26977f..0ffe508f3c704 100644 --- a/src/field_theory/normal.lean +++ b/src/field_theory/normal.lean @@ -138,9 +138,8 @@ local attribute [-instance] adjoin_root.has_smul lemma normal.of_is_splitting_field (p : F[X]) [hFEp : is_splitting_field F E p] : normal F E := begin unfreezingI { rcases eq_or_ne p 0 with rfl | hp }, - { have := hFEp.adjoin_roots, - simp only [polynomial.map_zero, roots_zero, multiset.to_finset_zero, finset.coe_empty, - algebra.adjoin_empty] at this, + { have := hFEp.adjoin_root_set, + simp only [root_set_zero, algebra.adjoin_empty] at this, exact normal.of_alg_equiv (alg_equiv.of_bijective (algebra.of_id F E) (algebra.bijective_algebra_map_iff.2 this.symm)) }, refine normal_iff.2 (λ x, _), @@ -197,7 +196,7 @@ begin dsimp only [S], rw [←finset.image_to_finset, finset.coe_image], apply eq.trans (algebra.adjoin_res_eq_adjoin_res F E C D - hFEp.adjoin_roots adjoin_root.adjoin_root_eq_top), + hFEp.adjoin_root_set adjoin_root.adjoin_root_eq_top), rw [set.image_singleton, ring_hom.algebra_map_to_algebra, adjoin_root.lift_root] end diff --git a/src/field_theory/polynomial_galois_group.lean b/src/field_theory/polynomial_galois_group.lean index 4c73c3d5bf348..ad3d3e5a0757e 100644 --- a/src/field_theory/polynomial_galois_group.lean +++ b/src/field_theory/polynomial_galois_group.lean @@ -63,7 +63,7 @@ alg_equiv.apply_mul_semiring_action begin refine alg_equiv.ext (λ x, (alg_hom.mem_equalizer σ.to_alg_hom τ.to_alg_hom x).mp ((set_like.ext_iff.mp _ x).mpr algebra.mem_top)), - rwa [eq_top_iff, ←splitting_field.adjoin_roots, algebra.adjoin_le_iff], + rwa [eq_top_iff, ←splitting_field.adjoin_root_set, algebra.adjoin_le_iff], end /-- If `p` splits in `F` then the `p.gal` is trivial. -/ diff --git a/src/field_theory/splitting_field/construction.lean b/src/field_theory/splitting_field/construction.lean index c160e90e30bc4..6236de333f50c 100644 --- a/src/field_theory/splitting_field/construction.lean +++ b/src/field_theory/splitting_field/construction.lean @@ -391,26 +391,28 @@ by { rw ← X_sub_C_mul_remove_factor _ hndf at hmf0, refine (splits_of_splits_m let ⟨k, hk⟩ := ih f.remove_factor (nat_degree_remove_factor' hf) (adjoin_root.lift j r hr) hsf in ⟨k, by rw [algebra_map_succ, ← ring_hom.comp_assoc, hk, adjoin_root.lift_comp_of]⟩ -theorem adjoin_roots (n : ℕ) : ∀ {K : Type u} [field K], by exactI +theorem adjoin_root_set (n : ℕ) : ∀ {K : Type u} [field K], by exactI ∀ (f : K[X]) (hfn : f.nat_degree = n), - algebra.adjoin K (↑(f.map $ algebra_map K $ splitting_field_aux n f).roots.to_finset : - set (splitting_field_aux n f)) = ⊤ := + algebra.adjoin K (f.root_set (splitting_field_aux n f)) = ⊤ := nat.rec_on n (λ K _ f hf, by exactI algebra.eq_top_iff.2 (λ x, subalgebra.range_le _ ⟨x, rfl⟩)) $ λ n ih K _ f hfn, by exactI have hndf : f.nat_degree ≠ 0, by { intro h, rw h at hfn, cases hfn }, have hfn0 : f ≠ 0, by { intro h, rw h at hndf, exact hndf rfl }, have hmf0 : map (algebra_map K (splitting_field_aux n.succ f)) f ≠ 0 := map_ne_zero hfn0, -by { rw [algebra_map_succ, ← map_map, ← X_sub_C_mul_remove_factor _ hndf, +begin + simp_rw root_set at ⊢ ih, + rw [algebra_map_succ, ← map_map, ← X_sub_C_mul_remove_factor _ hndf, polynomial.map_mul] at hmf0 ⊢, -rw [roots_mul hmf0, polynomial.map_sub, map_X, map_C, roots_X_sub_C, multiset.to_finset_add, - finset.coe_union, multiset.to_finset_singleton, finset.coe_singleton, - algebra.adjoin_union_eq_adjoin_adjoin, ← set.image_singleton, - algebra.adjoin_algebra_map K (adjoin_root f.factor) - (splitting_field_aux n f.remove_factor), - adjoin_root.adjoin_root_eq_top, algebra.map_top, - is_scalar_tower.adjoin_range_to_alg_hom K (adjoin_root f.factor) - (splitting_field_aux n f.remove_factor), - ih _ (nat_degree_remove_factor' hfn), subalgebra.restrict_scalars_top] } + rw [roots_mul hmf0, polynomial.map_sub, map_X, map_C, roots_X_sub_C, multiset.to_finset_add, + finset.coe_union, multiset.to_finset_singleton, finset.coe_singleton, + algebra.adjoin_union_eq_adjoin_adjoin, ← set.image_singleton, + algebra.adjoin_algebra_map K (adjoin_root f.factor) + (splitting_field_aux n f.remove_factor), + adjoin_root.adjoin_root_eq_top, algebra.map_top, + is_scalar_tower.adjoin_range_to_alg_hom K (adjoin_root f.factor) + (splitting_field_aux n f.remove_factor), + ih _ (nat_degree_remove_factor' hfn), subalgebra.restrict_scalars_top] +end end splitting_field_aux @@ -465,12 +467,8 @@ def lift : splitting_field f →ₐ[K] L := exact ring_hom.ext_iff.1 this r }, .. classical.some (splitting_field_aux.exists_lift _ _ _ _ hb) } -theorem adjoin_roots : algebra.adjoin K - (↑(f.map (algebra_map K $ splitting_field f)).roots.to_finset : set (splitting_field f)) = ⊤ := -splitting_field_aux.adjoin_roots _ _ rfl - -theorem adjoin_root_set : algebra.adjoin K (f.root_set f.splitting_field) = ⊤ := -adjoin_roots f +theorem adjoin_root_set : algebra.adjoin K (f.root_set (splitting_field f)) = ⊤ := +splitting_field_aux.adjoin_root_set _ _ rfl end splitting_field @@ -482,7 +480,7 @@ variables (K L) [algebra K L] variables {K} instance splitting_field (f : K[X]) : is_splitting_field K (splitting_field f) f := -⟨splitting_field.splits f, splitting_field.adjoin_roots f⟩ +⟨splitting_field.splits f, splitting_field.adjoin_root_set f⟩ instance (f : K[X]) : _root_.finite_dimensional K f.splitting_field := finite_dimensional f.splitting_field f diff --git a/src/field_theory/splitting_field/is_splitting_field.lean b/src/field_theory/splitting_field/is_splitting_field.lean index 8d93b01342067..1916a19ed67b9 100644 --- a/src/field_theory/splitting_field/is_splitting_field.lean +++ b/src/field_theory/splitting_field/is_splitting_field.lean @@ -45,7 +45,7 @@ variables [field K] [field L] [field F] [algebra K L] /-- Typeclass characterising splitting fields. -/ class is_splitting_field (f : K[X]) : Prop := (splits [] : splits (algebra_map K L) f) -(adjoin_roots [] : algebra.adjoin K (↑(f.map (algebra_map K L)).roots.to_finset : set L) = ⊤) +(adjoin_root_set [] : algebra.adjoin K (f.root_set L) = ⊤) variables {K L F} @@ -59,16 +59,17 @@ instance map (f : F[X]) [is_splitting_field F L f] : is_splitting_field K L (f.map $ algebra_map F K) := ⟨by { rw [splits_map_iff, ← is_scalar_tower.algebra_map_eq], exact splits L f }, subalgebra.restrict_scalars_injective F $ - by { rw [map_map, ← is_scalar_tower.algebra_map_eq, subalgebra.restrict_scalars_top, - eq_top_iff, ← adjoin_roots L f, algebra.adjoin_le_iff], + by { rw [root_set, map_map, ← is_scalar_tower.algebra_map_eq, subalgebra.restrict_scalars_top, + eq_top_iff, ← adjoin_root_set L f, algebra.adjoin_le_iff], exact λ x hx, @algebra.subset_adjoin K _ _ _ _ _ _ hx }⟩ variables (L) theorem splits_iff (f : K[X]) [is_splitting_field K L f] : polynomial.splits (ring_hom.id K) f ↔ (⊤ : subalgebra K L) = ⊥ := -⟨λ h, eq_bot_iff.2 $ adjoin_roots L f ▸ (roots_map (algebra_map K L) h).symm ▸ +⟨λ h, eq_bot_iff.2 $ adjoin_root_set L f ▸ algebra.adjoin_le_iff.2 (λ y hy, - let ⟨x, hxs, hxy⟩ := finset.mem_image.1 (by rwa multiset.to_finset_map at hy) in + let ⟨x, hxs, hxy⟩ := finset.mem_image.1 + (by rwa [root_set, roots_map _ h, multiset.to_finset_map] at hy) in hxy ▸ set_like.mem_coe.2 $ subalgebra.algebra_map_mem _ _), λ h, @ring_equiv.to_ring_hom_refl K _ ▸ ring_equiv.self_trans_symm (ring_equiv.of_bijective _ $ algebra.bijective_algebra_map_iff.2 h) ▸ @@ -80,14 +81,15 @@ theorem mul (f g : F[X]) (hf : f ≠ 0) (hg : g ≠ 0) [is_splitting_field F K f ⟨(is_scalar_tower.algebra_map_eq F K L).symm ▸ splits_mul _ (splits_comp_of_splits _ _ (splits K f)) ((splits_map_iff _ _).1 (splits L $ g.map $ algebra_map F K)), - by rw [polynomial.map_mul, roots_mul (mul_ne_zero (map_ne_zero hf : f.map (algebra_map F L) ≠ 0) + by rw [root_set, polynomial.map_mul, + roots_mul (mul_ne_zero (map_ne_zero hf : f.map (algebra_map F L) ≠ 0) (map_ne_zero hg)), multiset.to_finset_add, finset.coe_union, algebra.adjoin_union_eq_adjoin_adjoin, is_scalar_tower.algebra_map_eq F K L, ← map_map, roots_map (algebra_map K L) ((splits_id_iff_splits $ algebra_map F K).2 $ splits K f), - multiset.to_finset_map, finset.coe_image, algebra.adjoin_algebra_map, adjoin_roots, - algebra.map_top, is_scalar_tower.adjoin_range_to_alg_hom, ← map_map, adjoin_roots, - subalgebra.restrict_scalars_top]⟩ + multiset.to_finset_map, finset.coe_image, algebra.adjoin_algebra_map, ←root_set, + adjoin_root_set, algebra.map_top, is_scalar_tower.adjoin_range_to_alg_hom, ← map_map, + ←root_set, adjoin_root_set, subalgebra.restrict_scalars_top]⟩ end scalar_tower @@ -100,7 +102,7 @@ if hf0 : f = 0 then (algebra.of_id K F).comp $ (algebra.bot_equiv K L : (⊥ : subalgebra K L) →ₐ[K] K).comp $ by { rw ← (splits_iff L f).1 (show f.splits (ring_hom.id K), from hf0.symm ▸ splits_zero _), exact algebra.to_top } else -alg_hom.comp (by { rw ← adjoin_roots L f, exact classical.choice (lift_of_splits _ $ λ y hy, +alg_hom.comp (by { rw ← adjoin_root_set L f, exact classical.choice (lift_of_splits _ $ λ y hy, have aeval y f = 0, from (eval₂_eq_eval_map _).trans $ (mem_roots $ by exact map_ne_zero hf0).1 (multiset.mem_to_finset.mp hy), ⟨is_algebraic_iff_is_integral.1 ⟨f, hf0, this⟩, @@ -108,10 +110,10 @@ alg_hom.comp (by { rw ← adjoin_roots L f, exact classical.choice (lift_of_spli algebra.to_top theorem finite_dimensional (f : K[X]) [is_splitting_field K L f] : finite_dimensional K L := -⟨@algebra.top_to_submodule K L _ _ _ ▸ adjoin_roots L f ▸ +⟨@algebra.top_to_submodule K L _ _ _ ▸ adjoin_root_set L f ▸ fg_adjoin_of_finite (finset.finite_to_set _) (λ y hy, if hf : f = 0 - then by { rw [hf, polynomial.map_zero, roots_zero] at hy, cases hy } + then by { rw [hf, root_set_zero] at hy, cases hy } else is_algebraic_iff_is_integral.1 ⟨f, hf, (eval₂_eq_eval_map _).trans $ (mem_roots $ by exact map_ne_zero hf).1 (multiset.mem_to_finset.mp hy)⟩)⟩ @@ -122,7 +124,7 @@ begin { rw ← f.to_alg_hom.comp_algebra_map, exact splits_comp_of_splits _ _ (splits F p) }, { rw [←(algebra.range_top_iff_surjective f.to_alg_hom).mpr f.surjective, - ←root_set, adjoin_root_set_eq_range (splits F p), root_set, adjoin_roots F p] }, + adjoin_root_set_eq_range (splits F p), adjoin_root_set F p] }, end end is_splitting_field diff --git a/src/number_theory/cyclotomic/basic.lean b/src/number_theory/cyclotomic/basic.lean index 5a324524c74a3..da9fcb87c7960 100644 --- a/src/number_theory/cyclotomic/basic.lean +++ b/src/number_theory/cyclotomic/basic.lean @@ -363,29 +363,29 @@ section variables {A B} -lemma adjoin_roots_cyclotomic_eq_adjoin_nth_roots [decidable_eq B] [is_domain B] {ζ : B} +lemma adjoin_roots_cyclotomic_eq_adjoin_nth_roots [is_domain B] {ζ : B} {n : ℕ+} (hζ : is_primitive_root ζ n) : - adjoin A ↑((map (algebra_map A B) (cyclotomic n A)).roots.to_finset) = + adjoin A ((cyclotomic n A).root_set B) = adjoin A {b : B | ∃ (a : ℕ+), a ∈ ({n} : set ℕ+) ∧ b ^ (a : ℕ) = 1} := begin simp only [mem_singleton_iff, exists_eq_left, map_cyclotomic], refine le_antisymm (adjoin_mono (λ x hx, _)) (adjoin_le (λ x hx, _)), - { simp only [multiset.mem_to_finset, finset.mem_coe, - map_cyclotomic, mem_roots (cyclotomic_ne_zero n B)] at hx, + { rw [mem_root_set'] at hx, simp only [mem_singleton_iff, exists_eq_left, mem_set_of_eq], rw is_root_of_unity_iff n.pos, - exact ⟨n, nat.mem_divisors_self n n.ne_zero, hx⟩, - all_goals { apply_instance } }, + refine ⟨n, nat.mem_divisors_self n n.ne_zero, _⟩, + rw [is_root.def, ←map_cyclotomic n (algebra_map A B), eval_map, ←aeval_def], + exact hx.2 }, { simp only [mem_singleton_iff, exists_eq_left, mem_set_of_eq] at hx, obtain ⟨i, hin, rfl⟩ := hζ.eq_pow_of_pow_eq_one hx n.pos, refine set_like.mem_coe.2 (subalgebra.pow_mem _ (subset_adjoin _) _), - rwa [finset.mem_coe, multiset.mem_to_finset, mem_roots $ cyclotomic_ne_zero n B], - exact hζ.is_root_cyclotomic n.pos } + rw [mem_root_set', map_cyclotomic, aeval_def, ←eval_map, map_cyclotomic, ←is_root], + refine ⟨cyclotomic_ne_zero n B, hζ.is_root_cyclotomic n.pos⟩ } end -lemma adjoin_roots_cyclotomic_eq_adjoin_root_cyclotomic {n : ℕ+} [decidable_eq B] [is_domain B] +lemma adjoin_roots_cyclotomic_eq_adjoin_root_cyclotomic {n : ℕ+} [is_domain B] {ζ : B} (hζ : is_primitive_root ζ n) : - adjoin A (((map (algebra_map A B) (cyclotomic n A)).roots.to_finset) : set B) = adjoin A ({ζ}) := + adjoin A ((cyclotomic n A).root_set B) = adjoin A ({ζ}) := begin refine le_antisymm (adjoin_le (λ x hx, _)) (adjoin_mono (λ x hx, _)), { suffices hx : x ^ ↑n = 1, @@ -393,12 +393,11 @@ begin exact set_like.mem_coe.2 (subalgebra.pow_mem _ (subset_adjoin $ mem_singleton ζ) _), rw is_root_of_unity_iff n.pos, refine ⟨n, nat.mem_divisors_self n n.ne_zero, _⟩, - rwa [finset.mem_coe, multiset.mem_to_finset, - map_cyclotomic, mem_roots $ cyclotomic_ne_zero n B] at hx, - all_goals { apply_instance } }, + rw [mem_root_set', aeval_def, ←eval_map, map_cyclotomic, ←is_root] at hx, + exact hx.2 }, { simp only [mem_singleton_iff, exists_eq_left, mem_set_of_eq] at hx, - simpa only [hx, multiset.mem_to_finset, finset.mem_coe, map_cyclotomic, - mem_roots (cyclotomic_ne_zero n B)] using hζ.is_root_cyclotomic n.pos } + simpa only [hx, mem_root_set', map_cyclotomic, aeval_def, ←eval_map, is_root] + using and.intro (cyclotomic_ne_zero n B) (hζ.is_root_cyclotomic n.pos) } end lemma adjoin_primitive_root_eq_top {n : ℕ+} [is_domain B] [h : is_cyclotomic_extension {n} A B] @@ -468,15 +467,16 @@ variables [is_cyclotomic_extension {n} K L] /-- If `is_cyclotomic_extension {n} K L`, then `L` is the splitting field of `X ^ n - 1`. -/ lemma splitting_field_X_pow_sub_one : is_splitting_field K L (X ^ (n : ℕ) - 1) := { splits := splits_X_pow_sub_one K L (mem_singleton n), - adjoin_roots := + adjoin_root_set := begin rw [← ((iff_adjoin_eq_top {n} K L).1 infer_instance).2], congr, refine set.ext (λ x, _), simp only [polynomial.map_pow, mem_singleton_iff, multiset.mem_to_finset, exists_eq_left, mem_set_of_eq, polynomial.map_X, polynomial.map_one, finset.mem_coe, polynomial.map_sub], - rwa [← ring_hom.map_one C, mem_roots (@X_pow_sub_C_ne_zero L _ _ _ n.pos _), is_root.def, - eval_sub, eval_pow, eval_C, eval_X, sub_eq_zero] + simp only [mem_root_set', map_sub, map_pow, aeval_one, aeval_X, sub_eq_zero, map_X, + and_iff_right_iff_imp, polynomial.map_sub, polynomial.map_pow, polynomial.map_one], + exact λ _, X_pow_sub_C_ne_zero n.pos (1 : L) end } /-- Any two `n`-th cyclotomic extensions are isomorphic. -/ @@ -501,12 +501,14 @@ localized "attribute [instance] is_cyclotomic_extension.is_galois" in cyclotomic /-- If `is_cyclotomic_extension {n} K L`, then `L` is the splitting field of `cyclotomic n K`. -/ lemma splitting_field_cyclotomic : is_splitting_field K L (cyclotomic n K) := { splits := splits_cyclotomic K L (mem_singleton n), - adjoin_roots := + adjoin_root_set := begin rw [← ((iff_adjoin_eq_top {n} K L).1 infer_instance).2], letI := classical.dec_eq L, - obtain ⟨ζ, hζ⟩ := @is_cyclotomic_extension.exists_prim_root {n} K L _ _ _ _ _ (mem_singleton n), - exact adjoin_roots_cyclotomic_eq_adjoin_nth_roots hζ + -- todo: make `exists_prim_root` take an explicit `L` + obtain ⟨ζ : L, hζ⟩ := is_cyclotomic_extension.exists_prim_root K (mem_singleton n), + exact adjoin_roots_cyclotomic_eq_adjoin_nth_roots hζ, + all_goals { apply_instance } end } localized "attribute [instance] is_cyclotomic_extension.splitting_field_cyclotomic" in cyclotomic @@ -540,7 +542,7 @@ begin (splitting_field.splits _) (degree_cyclotomic_pos n K n.pos).ne', rw [← eval_map, ← is_root.def, map_cyclotomic, is_root_cyclotomic_iff] at hζ, refine ⟨forall_eq.2 ⟨ζ, hζ⟩, _⟩, - rw [←algebra.eq_top_iff, ←splitting_field.adjoin_roots, eq_comm], + rw [←algebra.eq_top_iff, ←splitting_field.adjoin_root_set, eq_comm], exact is_cyclotomic_extension.adjoin_roots_cyclotomic_eq_adjoin_nth_roots hζ, end @@ -668,7 +670,6 @@ instance [ne_zero ((n : ℕ) : A)] : lemma eq_adjoin_primitive_root {μ : (cyclotomic_field n K)} (h : is_primitive_root μ n) : cyclotomic_ring n A K = adjoin A ({μ} : set ((cyclotomic_field n K))) := begin - letI := classical.prop_decidable, rw [←is_cyclotomic_extension.adjoin_roots_cyclotomic_eq_adjoin_root_cyclotomic h, is_cyclotomic_extension.adjoin_roots_cyclotomic_eq_adjoin_nth_roots h], simp [cyclotomic_ring] From cec81510e48e579bde6acd8568c06a87af045b63 Mon Sep 17 00:00:00 2001 From: Amelia Livingston <101damnations@github.com> Date: Tue, 13 Jun 2023 01:21:10 +0000 Subject: [PATCH 1176/1291] refactor(representation_theory/Rep): define `ihom` concretely (#19170) --- src/representation_theory/Rep.lean | 204 ++++++++---------- .../group_cohomology/resolution.lean | 2 +- 2 files changed, 89 insertions(+), 117 deletions(-) diff --git a/src/representation_theory/Rep.lean b/src/representation_theory/Rep.lean index 01927f4e3b3dd..2f801e5652d25 100644 --- a/src/representation_theory/Rep.lean +++ b/src/representation_theory/Rep.lean @@ -68,7 +68,7 @@ lemma coe_of {V : Type u} [add_comm_group V] [module k V] (ρ : G →* (V →ₗ @[simp] lemma of_ρ {V : Type u} [add_comm_group V] [module k V] (ρ : G →* (V →ₗ[k] V)) : (of ρ).ρ = ρ := rfl -@[simp] lemma Action_ρ_eq_ρ {A : Rep k G} : Action.ρ A = A.ρ := rfl +lemma Action_ρ_eq_ρ {A : Rep k G} : Action.ρ A = A.ρ := rfl /-- Allows us to apply lemmas about the underlying `ρ`, which would take an element `g : G` rather than `g : Mon.of G` as an argument. -/ @@ -76,6 +76,18 @@ lemma of_ρ_apply {V : Type u} [add_comm_group V] [module k V] (ρ : representation k G V) (g : Mon.of G) : (Rep.of ρ).ρ g = ρ (g : G) := rfl +@[simp] lemma ρ_inv_self_apply {G : Type u} [group G] (A : Rep k G) (g : G) (x : A) : + A.ρ g⁻¹ (A.ρ g x) = x := +show (A.ρ g⁻¹ * A.ρ g) x = x, by rw [←map_mul, inv_mul_self, map_one, linear_map.one_apply] + +@[simp] lemma ρ_self_inv_apply {G : Type u} [group G] {A : Rep k G} (g : G) (x : A) : + A.ρ g (A.ρ g⁻¹ x) = x := +show (A.ρ g * A.ρ g⁻¹) x = x, by rw [←map_mul, mul_inv_self, map_one, linear_map.one_apply] + +lemma hom_comm_apply {A B : Rep k G} (f : A ⟶ B) (g : G) (x : A) : + f.hom (A.ρ g x) = B.ρ g (f.hom x) := +linear_map.ext_iff.1 (f.comm g) x + variables (k G) /-- The trivial `k`-linear `G`-representation on a `k`-module `V.` -/ @@ -234,151 +246,111 @@ end end linearization end -section +section monoidal_closed open Action variables [group G] (A B C : Rep k G) -noncomputable instance : monoidal_closed (Rep k G) := -monoidal_closed.of_equiv (functor_category_monoidal_equivalence _ _) - -/-- Explicit description of the 'internal Hom' `iHom(A, B)` of two representations `A, B`: -this is `F⁻¹(iHom(F(A), F(B)))`, where `F` is an equivalence -`Rep k G ≌ (single_obj G ⥤ Module k)`. Just used to prove `Rep.ihom_obj_ρ`. -/ -lemma ihom_obj_ρ_def : - ((ihom A).obj B).ρ = (functor_category_equivalence.inverse.obj - ((functor_category_equivalence.functor.obj A).closed_ihom.obj - (functor_category_equivalence.functor.obj B))).ρ := rfl - -/-- Given `k`-linear `G`-representations `(A, ρ₁), (B, ρ₂)`, the 'internal Hom' is the -representation on `Homₖ(A, B)` sending `g : G` and `f : A →ₗ[k] B` to `(ρ₂ g) ∘ₗ f ∘ₗ (ρ₁ g⁻¹)`. -/ -@[simp] lemma ihom_obj_ρ : - ((ihom A).obj B).ρ = A.ρ.lin_hom B.ρ := -begin - refine monoid_hom.ext (λ g, _), - simpa only [ihom_obj_ρ_def, functor_category_equivalence.inverse_obj_ρ_apply, - functor.closed_ihom_obj_map, ←functor.map_inv, single_obj.inv_as_inv], -end - -@[simp] lemma ihom_map_hom {B C : Rep k G} (f : B ⟶ C) : - ((ihom A).map f).hom = linear_map.llcomp k A B C f.hom := -rfl - -/-- Unfolds the unit in the adjunction `A ⊗ - ⊣ iHom(A, -)`; just used to prove -`Rep.ihom_coev_app_hom`. -/ -lemma ihom_coev_app_def : - (ihom.coev A).app B = functor_category_equivalence.unit_iso.hom.app B ≫ - functor_category_equivalence.inverse.map - ((functor_category_equivalence.functor.obj A).closed_unit.app _ ≫ - (functor_category_equivalence.functor.obj A).closed_ihom.map - ((functor_category_monoidal_equivalence (Module.{u} k) (Mon.of G)).μ A B)) := -rfl - -/-- Describes the unit in the adjunction `A ⊗ - ⊣ iHom(A, -)`; given another `k`-linear -`G`-representation `B,` the `k`-linear map underlying the resulting map `B ⟶ iHom(A, A ⊗ B)` is -given by flipping the arguments in the natural `k`-bilinear map `A →ₗ[k] B →ₗ[k] A ⊗ B`. -/ -@[simp] lemma ihom_coev_app_hom : - Action.hom.hom ((ihom.coev A).app B) = (tensor_product.mk _ _ _).flip := -begin - refine linear_map.ext (λ x, linear_map.ext (λ y, _)), - simpa only [ihom_coev_app_def, functor.map_comp, comp_hom, - functor_category_equivalence.inverse_map_hom, functor.closed_ihom_map_app, - functor_category_monoidal_equivalence.μ_app], -end - -variables {A B C} - -/-- Given a `k`-linear `G`-representation `A`, the adjunction `A ⊗ - ⊣ iHom(A, -)` defines a -bijection `Hom(A ⊗ B, C) ≃ Hom(B, iHom(A, C))` for all `B, C`. Given `f : A ⊗ B ⟶ C`, this lemma -describes the `k`-linear map underlying `f`'s image under the bijection. It is given by currying the +/-- Given a `k`-linear `G`-representation `(A, ρ₁)`, this is the 'internal Hom' functor sending +`(B, ρ₂)` to the representation `Homₖ(A, B)` that maps `g : G` and `f : A →ₗ[k] B` to +`(ρ₂ g) ∘ₗ f ∘ₗ (ρ₁ g⁻¹)`. -/ +@[simps] protected def ihom (A : Rep k G) : Rep k G ⥤ Rep k G := +{ obj := λ B, Rep.of (representation.lin_hom A.ρ B.ρ), + map := λ X Y f, + { hom := Module.of_hom (linear_map.llcomp k _ _ _ f.hom), + comm' := λ g, linear_map.ext (λ x, linear_map.ext (λ y, + show f.hom (X.ρ g _) = _, by simpa only [hom_comm_apply])) }, + map_id' := λ B, by ext; refl, + map_comp' := λ B C D f g, by ext; refl } + +@[simp] protected lemma ihom_obj_ρ_apply {A B : Rep k G} (g : G) (x : A →ₗ[k] B) : + ((Rep.ihom A).obj B).ρ g x = (B.ρ g) ∘ₗ x ∘ₗ (A.ρ g⁻¹) := rfl + +/-- Given a `k`-linear `G`-representation `A`, this is the Hom-set bijection in the adjunction +`A ⊗ - ⊣ ihom(A, -)`. It sends `f : A ⊗ B ⟶ C` to a `Rep k G` morphism defined by currying the `k`-linear map underlying `f`, giving a map `A →ₗ[k] B →ₗ[k] C`, then flipping the arguments. -/ -@[simp] lemma monoidal_closed_curry_hom (f : A ⊗ B ⟶ C) : - (monoidal_closed.curry f).hom = (tensor_product.curry f.hom).flip := -begin - rw [monoidal_closed.curry_eq, comp_hom, ihom_coev_app_hom], - refl, -end - -/-- Given a `k`-linear `G`-representation `A`, the adjunction `A ⊗ - ⊣ iHom(A, -)` defines a -bijection `Hom(A ⊗ B, C) ≃ Hom(B, iHom(A, C))` for all `B, C`. Given `f : B ⟶ iHom(A, C)`, this -lemma describes the `k`-linear map underlying `f`'s image under the bijection. It is given by -flipping the arguments of the `k`-linear map underlying `f`, giving a map `A →ₗ[k] B →ₗ[k] C`, then -uncurrying. -/ -@[simp] lemma monoidal_closed_uncurry_hom (f : B ⟶ (ihom A).obj C) : - (monoidal_closed.uncurry f).hom = tensor_product.uncurry _ _ _ _ f.hom.flip := -begin - simp only [monoidal_closed.of_equiv_uncurry_def, comp_inv_iso_inv_app, - monoidal_functor.comm_tensor_left_inv_app, comp_hom, - functor_category_monoidal_equivalence.inverse_map, functor_category_equivalence.inverse_map_hom, - functor_category_monoidal_equivalence.μ_iso_inv_app], - ext, - refl, -end - -/-- Describes the counit in the adjunction `A ⊗ - ⊣ iHom(A, -)`; given another `k`-linear -`G`-representation `B,` the `k`-linear map underlying the resulting morphism `A ⊗ iHom(A, B) ⟶ B` -is given by uncurrying the map `A →ₗ[k] (A →ₗ[k] B) →ₗ[k] B` defined by flipping the arguments in -the identity map on `Homₖ(A, B).` -/ -@[simp] lemma ihom_ev_app_hom : Action.hom.hom ((ihom.ev A).app B) = - tensor_product.uncurry _ _ _ _ linear_map.id.flip := -monoidal_closed_uncurry_hom _ +@[simps] protected def hom_equiv (A B C : Rep k G) : (A ⊗ B ⟶ C) ≃ (B ⟶ (Rep.ihom A).obj C) := +{ to_fun := λ f, + { hom := (tensor_product.curry f.hom).flip, + comm' := λ g, + begin + refine linear_map.ext (λ x, linear_map.ext (λ y, _)), + change f.hom (_ ⊗ₜ[k] _) = C.ρ g (f.hom (_ ⊗ₜ[k] _)), + rw [←hom_comm_apply], + change _ = f.hom ((A.ρ g * A.ρ g⁻¹) y ⊗ₜ[k] _), + simpa only [←map_mul, mul_inv_self, map_one], + end }, + inv_fun := λ f, + { hom := tensor_product.uncurry k _ _ _ f.hom.flip, + comm' := λ g, tensor_product.ext' (λ x y, + begin + dsimp only [monoidal_category.tensor_left_obj, Module.comp_def, linear_map.comp_apply, + tensor_rho, Module.monoidal_category.hom_apply, tensor_product.map_tmul], + simp only [tensor_product.uncurry_apply f.hom.flip, linear_map.flip_apply, + Action_ρ_eq_ρ, hom_comm_apply f g y, Rep.ihom_obj_ρ_apply, linear_map.comp_apply, + ρ_inv_self_apply], + end) }, + left_inv := λ f, Action.hom.ext _ _ (tensor_product.ext' (λ x y, rfl)), + right_inv := λ f, by ext; refl } + +instance : monoidal_closed (Rep k G) := +{ closed' := λ A, + { is_adj := + { right := Rep.ihom A, + adj := adjunction.mk_of_hom_equiv + { hom_equiv := Rep.hom_equiv A, + hom_equiv_naturality_left_symm' := λ X Y Z f g, by ext; refl, + hom_equiv_naturality_right' := λ X Y Z f g, by ext; refl } } } } + +@[simp] lemma ihom_obj_ρ_def (A B : Rep k G) : ((ihom A).obj B).ρ = ((Rep.ihom A).obj B).ρ := rfl + +@[simp] lemma hom_equiv_def (A B C : Rep k G) : + (ihom.adjunction A).hom_equiv B C = Rep.hom_equiv A B C := rfl + +@[simp] lemma ihom_ev_app_hom (A B : Rep k G) : + Action.hom.hom ((ihom.ev A).app B) + = tensor_product.uncurry _ _ _ _ linear_map.id.flip := +by ext; refl + +@[simp] lemma ihom_coev_app_hom (A B : Rep k G) : + Action.hom.hom ((ihom.coev A).app B) = (tensor_product.mk _ _ _).flip := +by ext; refl variables (A B C) /-- There is a `k`-linear isomorphism between the sets of representation morphisms`Hom(A ⊗ B, C)` and `Hom(B, Homₖ(A, C))`. -/ -noncomputable def monoidal_closed.linear_hom_equiv : +def monoidal_closed.linear_hom_equiv : (A ⊗ B ⟶ C) ≃ₗ[k] (B ⟶ (A ⟶[Rep k G] C)) := { map_add' := λ f g, rfl, map_smul' := λ r f, rfl, ..(ihom.adjunction A).hom_equiv _ _ } /-- There is a `k`-linear isomorphism between the sets of representation morphisms`Hom(A ⊗ B, C)` and `Hom(A, Homₖ(B, C))`. -/ -noncomputable def monoidal_closed.linear_hom_equiv_comm : +def monoidal_closed.linear_hom_equiv_comm : (A ⊗ B ⟶ C) ≃ₗ[k] (A ⟶ (B ⟶[Rep k G] C)) := (linear.hom_congr k (β_ A B) (iso.refl _)) ≪≫ₗ monoidal_closed.linear_hom_equiv _ _ _ variables {A B C} -lemma monoidal_closed.linear_hom_equiv_hom (f : A ⊗ B ⟶ C) : +@[simp] lemma monoidal_closed.linear_hom_equiv_hom (f : A ⊗ B ⟶ C) : (monoidal_closed.linear_hom_equiv A B C f).hom = - (tensor_product.curry f.hom).flip := -monoidal_closed_curry_hom _ + (tensor_product.curry f.hom).flip := rfl -lemma monoidal_closed.linear_hom_equiv_comm_hom (f : A ⊗ B ⟶ C) : +@[simp] lemma monoidal_closed.linear_hom_equiv_comm_hom (f : A ⊗ B ⟶ C) : (monoidal_closed.linear_hom_equiv_comm A B C f).hom = - tensor_product.curry f.hom := -begin - dunfold monoidal_closed.linear_hom_equiv_comm, - refine linear_map.ext (λ x, linear_map.ext (λ y, _)), - simp only [linear_equiv.trans_apply, monoidal_closed.linear_hom_equiv_hom, - linear.hom_congr_apply, iso.refl_hom, iso.symm_hom, linear_map.to_fun_eq_coe, - linear_map.coe_comp, function.comp_app, linear.left_comp_apply, linear.right_comp_apply, - category.comp_id, Action.comp_hom, linear_map.flip_apply, tensor_product.curry_apply, - Module.coe_comp, function.comp_app, monoidal_category.braiding_inv_apply], -end + tensor_product.curry f.hom := rfl -lemma monoidal_closed.linear_hom_equiv_symm_hom (f : B ⟶ (A ⟶[Rep k G] C)) : +@[simp] lemma monoidal_closed.linear_hom_equiv_symm_hom (f : B ⟶ (A ⟶[Rep k G] C)) : ((monoidal_closed.linear_hom_equiv A B C).symm f).hom = - tensor_product.uncurry k A B C f.hom.flip := -monoidal_closed_uncurry_hom _ + tensor_product.uncurry k A B C f.hom.flip := rfl -lemma monoidal_closed.linear_hom_equiv_comm_symm_hom (f : A ⟶ (B ⟶[Rep k G] C)) : +@[simp] lemma monoidal_closed.linear_hom_equiv_comm_symm_hom (f : A ⟶ (B ⟶[Rep k G] C)) : ((monoidal_closed.linear_hom_equiv_comm A B C).symm f).hom = tensor_product.uncurry k A B C f.hom := -begin - dunfold monoidal_closed.linear_hom_equiv_comm, - refine tensor_product.algebra_tensor_module.curry_injective - (linear_map.ext (λ x, linear_map.ext (λ y, _))), - simp only [linear_equiv.trans_symm, linear_equiv.trans_apply, linear.hom_congr_symm_apply, - iso.refl_inv, linear_map.coe_comp, function.comp_app, category.comp_id, Action.comp_hom, - monoidal_closed.linear_hom_equiv_symm_hom, tensor_product.algebra_tensor_module.curry_apply, - linear_map.coe_restrict_scalars, linear_map.to_fun_eq_coe, linear_map.flip_apply, - tensor_product.curry_apply, Module.coe_comp, function.comp_app, - monoidal_category.braiding_hom_apply, tensor_product.uncurry_apply], -end +by ext; refl -end +end monoidal_closed end Rep namespace representation variables {k G : Type u} [comm_ring k] [monoid G] {V W : Type u} diff --git a/src/representation_theory/group_cohomology/resolution.lean b/src/representation_theory/group_cohomology/resolution.lean index 1d63d4ee38387..434f70e9d0ab5 100644 --- a/src/representation_theory/group_cohomology/resolution.lean +++ b/src/representation_theory/group_cohomology/resolution.lean @@ -276,7 +276,7 @@ begin iso.refl_inv, category.comp_id, Rep.monoidal_closed.linear_hom_equiv_comm_symm_hom, iso.trans_hom, Module.comp_def, linear_map.comp_apply, representation.Rep_of_tprod_iso_apply, diagonal_succ_hom_single x (1 : k), tensor_product.uncurry_apply, Rep.left_regular_hom_hom, - finsupp.lift_apply, Rep.ihom_obj_ρ, representation.lin_hom_apply, finsupp.sum_single_index, + finsupp.lift_apply, ihom_obj_ρ_def, Rep.ihom_obj_ρ_apply, finsupp.sum_single_index, zero_smul, one_smul, Rep.of_ρ, Rep.Action_ρ_eq_ρ, Rep.trivial_def (x 0)⁻¹, finsupp.llift_apply A k k], end From 44e2ae8cffc713925494e4975ee31ec1d06929b3 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Tue, 13 Jun 2023 06:00:19 +0000 Subject: [PATCH 1177/1291] chore(*): add mathlib4 synchronization comments (#19181) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.category.Module.simple` * `algebra.category.Module.subobject` * `algebra.lie.engel` * `analysis.normed_space.star.continuous_functional_calculus` * `category_theory.bicategory.functor_bicategory` * `category_theory.groupoid.subgroupoid` * `category_theory.monoidal.center` * `linear_algebra.eigenspace.minpoly` * `measure_theory.integral.torus_integral` * `probability.density` * `topology.metric_space.hausdorff_dimension` --- src/algebra/category/Module/simple.lean | 3 +++ src/algebra/category/Module/subobject.lean | 3 +++ src/algebra/lie/engel.lean | 3 +++ .../normed_space/star/continuous_functional_calculus.lean | 3 +++ src/category_theory/bicategory/functor_bicategory.lean | 3 +++ src/category_theory/groupoid/subgroupoid.lean | 3 +++ src/category_theory/monoidal/center.lean | 3 +++ src/linear_algebra/eigenspace/minpoly.lean | 3 +++ src/measure_theory/integral/torus_integral.lean | 3 +++ src/probability/density.lean | 3 +++ src/topology/metric_space/hausdorff_dimension.lean | 3 +++ 11 files changed, 33 insertions(+) diff --git a/src/algebra/category/Module/simple.lean b/src/algebra/category/Module/simple.lean index 4d2453c3be8a3..1ed951e27ab24 100644 --- a/src/algebra/category/Module/simple.lean +++ b/src/algebra/category/Module/simple.lean @@ -12,6 +12,9 @@ import linear_algebra.finite_dimensional /-! # Simple objects in the category of `R`-modules +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We prove simple modules are exactly simple objects in the category of `R`-modules. -/ diff --git a/src/algebra/category/Module/subobject.lean b/src/algebra/category/Module/subobject.lean index ea03103006f3c..38dd5733b0a9f 100644 --- a/src/algebra/category/Module/subobject.lean +++ b/src/algebra/category/Module/subobject.lean @@ -11,6 +11,9 @@ import category_theory.subobject.limits /-! # Subobjects in the category of `R`-modules +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We construct an explicit order isomorphism between the categorical subobjects of an `R`-module `M` and its submodules. This immediately implies that the category of `R`-modules is well-powered. diff --git a/src/algebra/lie/engel.lean b/src/algebra/lie/engel.lean index 23045db7fcb71..9e22f226cc1fa 100644 --- a/src/algebra/lie/engel.lean +++ b/src/algebra/lie/engel.lean @@ -9,6 +9,9 @@ import algebra.lie.normalizer /-! # Engel's theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains a proof of Engel's theorem providing necessary and sufficient conditions for Lie algebras and Lie modules to be nilpotent. diff --git a/src/analysis/normed_space/star/continuous_functional_calculus.lean b/src/analysis/normed_space/star/continuous_functional_calculus.lean index f3abf11cc57f4..98f628a932ced 100644 --- a/src/analysis/normed_space/star/continuous_functional_calculus.lean +++ b/src/analysis/normed_space/star/continuous_functional_calculus.lean @@ -9,6 +9,9 @@ import topology.algebra.star_subalgebra /-! # Continuous functional calculus +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we construct the `continuous_functional_calculus` for a normal element `a` of a (unital) C⋆-algebra over `ℂ`. This is a star algebra equivalence `C(spectrum ℂ a, ℂ) ≃⋆ₐ[ℂ] elemental_star_algebra ℂ a` which sends the (restriction of) the diff --git a/src/category_theory/bicategory/functor_bicategory.lean b/src/category_theory/bicategory/functor_bicategory.lean index 0d1cd9f0e4618..a626b310d3ced 100644 --- a/src/category_theory/bicategory/functor_bicategory.lean +++ b/src/category_theory/bicategory/functor_bicategory.lean @@ -8,6 +8,9 @@ import category_theory.bicategory.natural_transformation /-! # The bicategory of oplax functors between two bicategories +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given bicategories `B` and `C`, we give a bicategory structure on `oplax_functor B C` whose * objects are oplax functors, * 1-morphisms are oplax natural transformations, and diff --git a/src/category_theory/groupoid/subgroupoid.lean b/src/category_theory/groupoid/subgroupoid.lean index 04d52184a7323..e605586966081 100644 --- a/src/category_theory/groupoid/subgroupoid.lean +++ b/src/category_theory/groupoid/subgroupoid.lean @@ -13,6 +13,9 @@ import order.galois_connection /-! # Subgroupoid +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines subgroupoids as `structure`s containing the subsets of arrows and their stability under composition and inversion. Also defined are: diff --git a/src/category_theory/monoidal/center.lean b/src/category_theory/monoidal/center.lean index 4c6eb069a90e3..3d109154edcee 100644 --- a/src/category_theory/monoidal/center.lean +++ b/src/category_theory/monoidal/center.lean @@ -10,6 +10,9 @@ import category_theory.monoidal.coherence /-! # Half braidings and the Drinfeld center of a monoidal category +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define `center C` to be pairs `⟨X, b⟩`, where `X : C` and `b` is a half-braiding on `X`. We show that `center C` is braided monoidal, diff --git a/src/linear_algebra/eigenspace/minpoly.lean b/src/linear_algebra/eigenspace/minpoly.lean index feafdd2ee815f..f371c2eadd8ab 100644 --- a/src/linear_algebra/eigenspace/minpoly.lean +++ b/src/linear_algebra/eigenspace/minpoly.lean @@ -10,6 +10,9 @@ import field_theory.minpoly.field /-! # Eigenvalues are the roots of the minimal polynomial. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Tags eigenvalue, minimal polynomial diff --git a/src/measure_theory/integral/torus_integral.lean b/src/measure_theory/integral/torus_integral.lean index 02151c688526a..57bd43db674a7 100644 --- a/src/measure_theory/integral/torus_integral.lean +++ b/src/measure_theory/integral/torus_integral.lean @@ -9,6 +9,9 @@ import measure_theory.integral.circle_integral /-! # Integral over a torus in `ℂⁿ` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + 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}$, diff --git a/src/probability/density.lean b/src/probability/density.lean index e4c992f80cc1c..a4c7666e7d7e5 100644 --- a/src/probability/density.lean +++ b/src/probability/density.lean @@ -9,6 +9,9 @@ import measure_theory.measure.haar.of_basis /-! # Probability density function +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the probability density function of random variables, by which we mean measurable functions taking values in a Borel space. In particular, a measurable function `f` is said to the probability density function of a random variable `X` if for all measurable diff --git a/src/topology/metric_space/hausdorff_dimension.lean b/src/topology/metric_space/hausdorff_dimension.lean index 2dee87d9172dd..fbad19511cd6d 100644 --- a/src/topology/metric_space/hausdorff_dimension.lean +++ b/src/topology/metric_space/hausdorff_dimension.lean @@ -9,6 +9,9 @@ import measure_theory.measure.hausdorff /-! # Hausdorff dimension +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The Hausdorff dimension of a set `X` in an (extended) metric space is the unique number `dimH s : ℝ≥0∞` such that for any `d : ℝ≥0` we have From e2e7f2ac359e7514e4d40061d7c08bb69487ba4e Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Tue, 13 Jun 2023 07:06:56 +0000 Subject: [PATCH 1178/1291] feat(algebraic_geometry/elliptic_curve/point): the group law on the nonsingular rational points of a Weierstrass curve (#18000) Define a group structure on the nonsingular rational points of a Weierstrass curve by constructing an explicit injective group homomorphism into the class group of its coordinate ring and therefore proving associativity of the addition law. Co-authored-by: Junyan Xu - [x] depends on: #17999 - [x] depends on: #17194 - [x] depends on: #18038 - [x] depends on: #18101 - [x] depends on: #19121 Co-authored-by: Junyan Xu --- .../elliptic_curve/point.lean | 283 +++++++++++++++++- .../elliptic_curve/weierstrass.lean | 41 ++- 2 files changed, 298 insertions(+), 26 deletions(-) diff --git a/src/algebraic_geometry/elliptic_curve/point.lean b/src/algebraic_geometry/elliptic_curve/point.lean index 59c6aa82ed658..e347de9ad650d 100644 --- a/src/algebraic_geometry/elliptic_curve/point.lean +++ b/src/algebraic_geometry/elliptic_curve/point.lean @@ -5,20 +5,22 @@ Authors: David Kurniadi Angdinata -/ import algebraic_geometry.elliptic_curve.weierstrass +import linear_algebra.free_module.norm +import ring_theory.class_group /-! -# The group of nonsingular rational points on a Weierstrass curve over a field +# Nonsingular rational points on Weierstrass curves This file defines the type of nonsingular rational points on a Weierstrass curve over a field and -(TODO) proves that it forms an abelian group under a geometric secant-and-tangent process. +proves that it forms an abelian group under a geometric secant-and-tangent process. ## Mathematical background Let `W` be a Weierstrass curve over a field `F`. A rational point on `W` is simply a point -$[A:B:C]$ defined over `F` in the projective plane satisfying the homogeneous cubic equation -$B^2C + a_1ABC + a_3BC^2 = A^3 + a_2A^2C + a_4AC^2 + a_6C^3$. Any such point either lies in the -affine chart $C \ne 0$ and satisfies the Weierstrass equation obtained by setting $X := A/C$ and -$Y := B/C$, or is the unique point at infinity $0 := [0:1:0]$ when $C = 0$. With this new +$[X:Y:Z]$ defined over `F` in the projective plane satisfying the homogeneous cubic equation +$Y^2Z + a_1XYZ + a_3YZ^2 = X^3 + a_2X^2Z + a_4XZ^2 + a_6Z^3$. Any such point either lies in the +affine chart $Z \ne 0$ and satisfies the Weierstrass equation obtained by replacing $X/Z$ with $X$ +and $Y/Z$ with $Y$, or is the unique point at infinity $0 := [0:1:0]$ when $Z = 0$. With this new description, a nonsingular rational point on `W` is either $0$ or an affine point $(x, y)$ where the partial derivatives $W_X(X, Y)$ and $W_Y(X, Y)$ do not vanish simultaneously. For a field extension `K` of `F`, a `K`-rational point is simply a rational point on `W` base changed to `K`. @@ -50,11 +52,12 @@ The group law on this set is then uniquely determined by these constructions. ## Main statements - * TODO: the addition of two nonsingular rational points on `W` forms a group. + * `weierstrass_curve.point.add_comm_group`: the type of nonsingular rational points on `W` forms an + abelian group under addition. ## Notations - * `W⟮K⟯`: the group of nonsingular rational points on a Weierstrass curve `W` base changed to `K`. + * `W⟮K⟯`: the group of nonsingular rational points on `W` base changed to `K`. ## References @@ -66,7 +69,8 @@ elliptic curve, rational point, group law -/ private meta def map_simp : tactic unit := -`[simp only [map_one, map_bit0, map_bit1, map_neg, map_add, map_sub, map_mul, map_pow, map_div₀]] +`[simp only [map_one, map_bit0, map_bit1, map_neg, map_add, map_sub, _root_.map_mul, map_pow, + map_div₀]] private meta def eval_simp : tactic unit := `[simp only [eval_C, eval_X, eval_neg, eval_add, eval_sub, eval_mul, eval_pow]] @@ -82,9 +86,9 @@ universes u v w namespace weierstrass_curve -open polynomial +open coordinate_ring ideal polynomial -open_locale polynomial polynomial_polynomial +open_locale non_zero_divisors polynomial polynomial_polynomial section basic @@ -123,7 +127,16 @@ with a slope of $L$ that passes through an affine point $(x_1, y_1)$. This does not depend on `W`, and has argument order: $x_1$, $y_1$, $L$. -/ noncomputable def line_polynomial : R[X] := C L * (X - C x₁) + C y₁ -/-- The polynomial obtained by substituting the line $Y = L(X - x_1) + y_1$, with a slope of $L$ +lemma XY_ideal_eq₁ : XY_ideal W x₁ (C y₁) = XY_ideal W x₁ (line_polynomial x₁ y₁ L) := +begin + simp only [XY_ideal, X_class, Y_class, line_polynomial], + rw [← span_pair_add_mul_right $ adjoin_root.mk _ $ C $ C $ -L, ← _root_.map_mul, ← map_add], + apply congr_arg (_ ∘ _ ∘ _ ∘ _), + C_simp, + ring1 +end + +/-- The polynomial obtained by substituting the line $Y = L*(X - x_1) + y_1$, with a slope of $L$ that passes through an affine point $(x_1, y_1)$, into the polynomial $W(X, Y)$ associated to `W`. If such a line intersects `W` at another point $(x_2, y_2)$, then the roots of this polynomial are precisely $x_1$, $x_2$, and the $X$-coordinate of the addition of $(x_1, y_1)$ and $(x_2, y_2)$. @@ -131,6 +144,19 @@ precisely $x_1$, $x_2$, and the $X$-coordinate of the addition of $(x_1, y_1)$ a This depends on `W`, and has argument order: $x_1$, $y_1$, $L$. -/ noncomputable def add_polynomial : R[X] := W.polynomial.eval $ line_polynomial x₁ y₁ L +lemma C_add_polynomial : + C (W.add_polynomial x₁ y₁ L) + = (Y - C (line_polynomial x₁ y₁ L)) * (W.neg_polynomial - C (line_polynomial x₁ y₁ L)) + + W.polynomial := +by { rw [add_polynomial, line_polynomial, weierstrass_curve.polynomial, neg_polynomial], eval_simp, + C_simp, ring1 } + +lemma coordinate_ring.C_add_polynomial : + adjoin_root.mk W.polynomial (C (W.add_polynomial x₁ y₁ L)) + = adjoin_root.mk W.polynomial + ((Y - C (line_polynomial x₁ y₁ L)) * (W.neg_polynomial - C (line_polynomial x₁ y₁ L))) := +adjoin_root.mk_eq_mk.mpr ⟨1, by rw [C_add_polynomial, add_sub_cancel', mul_one]⟩ + lemma add_polynomial_eq : W.add_polynomial x₁ y₁ L = -cubic.to_poly ⟨1, -L ^ 2 - W.a₁ * L + W.a₂, 2 * x₁ * L ^ 2 + (W.a₁ * x₁ - 2 * y₁ - W.a₃) * L + (-W.a₁ * y₁ + W.a₄), @@ -186,6 +212,20 @@ lemma base_change_add_Y_of_base_change (x₁ x₂ y₁ L : A) : (algebra_map A B L) = algebra_map A B ((W.base_change A).add_Y x₁ x₂ y₁ L) := by rw [← base_change_add_Y, base_change_base_change] +lemma XY_ideal_add_eq : + XY_ideal W (W.add_X x₁ x₂ L) (C (W.add_Y x₁ x₂ y₁ L)) + = span {adjoin_root.mk W.polynomial $ W.neg_polynomial - C (line_polynomial x₁ y₁ L)} + ⊔ X_ideal W (W.add_X x₁ x₂ L) := +begin + simp only [XY_ideal, X_ideal, X_class, Y_class, add_Y, add_Y', neg_Y, neg_polynomial, + line_polynomial], + conv_rhs { rw [sub_sub, ← neg_add', map_neg, span_singleton_neg, sup_comm, ← span_insert] }, + rw [← span_pair_add_mul_right $ adjoin_root.mk _ $ C $ C $ W.a₁ + L, ← _root_.map_mul, ← map_add], + apply congr_arg (_ ∘ _ ∘ _ ∘ _), + C_simp, + ring1 +end + lemma equation_add_iff : W.equation (W.add_X x₁ x₂ L) (W.add_Y' x₁ x₂ y₁ L) ↔ (W.add_polynomial x₁ y₁ L).eval (W.add_X x₁ x₂ L) = 0 := @@ -359,6 +399,27 @@ end lemma Y_eq_of_Y_ne (hx : x₁ = x₂) (hy : y₁ ≠ W.neg_Y x₂ y₂) : y₁ = y₂ := or.resolve_right (Y_eq_of_X_eq h₁' h₂' hx) hy +lemma XY_ideal_eq₂ (hxy : x₁ = x₂ → y₁ ≠ W.neg_Y x₂ y₂) : + XY_ideal W x₂ (C y₂) = XY_ideal W x₂ (line_polynomial x₁ y₁ $ W.slope x₁ x₂ y₁ y₂) := +begin + have hy₂ : y₂ = (line_polynomial x₁ y₁ $ W.slope x₁ x₂ y₁ y₂).eval x₂ := + begin + by_cases hx : x₁ = x₂, + { rcases ⟨hx, Y_eq_of_Y_ne h₁' h₂' hx $ hxy hx⟩ with ⟨rfl, rfl⟩, + field_simp [line_polynomial, sub_ne_zero_of_ne (hxy rfl)] }, + { field_simp [line_polynomial, slope_of_X_ne hx, sub_ne_zero_of_ne hx], + ring1 } + end, + nth_rewrite_lhs 0 [hy₂], + simp only [XY_ideal, X_class, Y_class, line_polynomial], + rw [← span_pair_add_mul_right $ adjoin_root.mk W.polynomial $ C $ C $ -W.slope x₁ x₂ y₁ y₂, + ← _root_.map_mul, ← map_add], + apply congr_arg (_ ∘ _ ∘ _ ∘ _), + eval_simp, + C_simp, + ring1 +end + lemma add_polynomial_slope (hxy : x₁ = x₂ → y₁ ≠ W.neg_Y x₂ y₂) : W.add_polynomial x₁ y₁ (W.slope x₁ x₂ y₁ y₂) = -((X - C x₁) * (X - C x₂) * (X - C (W.add_X x₁ x₂ $ W.slope x₁ x₂ y₁ y₂))) := @@ -390,6 +451,11 @@ begin with { normalization_tactic := `[field_simp [hx], ring1] } } } end +lemma coordinate_ring.C_add_polynomial_slope (hxy : x₁ = x₂ → y₁ ≠ W.neg_Y x₂ y₂) : + adjoin_root.mk W.polynomial (C $ W.add_polynomial x₁ y₁ $ W.slope x₁ x₂ y₁ y₂) + = -(X_class W x₁ * X_class W x₂ * X_class W (W.add_X x₁ x₂ $ W.slope x₁ x₂ y₁ y₂)) := +by simpa only [add_polynomial_slope h₁' h₂' hxy, map_neg, neg_inj, _root_.map_mul] + lemma derivative_add_polynomial_slope (hxy : x₁ = x₂ → y₁ ≠ W.neg_Y x₂ y₂) : derivative (W.add_polynomial x₁ y₁ $ W.slope x₁ x₂ y₁ y₂) = -((X - C x₁) * (X - C x₂) + (X - C x₁) * (X - C (W.add_X x₁ x₂ $ W.slope x₁ x₂ y₁ y₂)) @@ -502,15 +568,168 @@ section group /-! ### The axioms for nonsingular rational points on a Weierstrass curve -/ -variables {F : Type u} [field F] {W : weierstrass_curve F} +variables {F : Type u} [field F] {W : weierstrass_curve F} {x₁ x₂ y₁ y₂ : F} + (h₁ : W.nonsingular x₁ y₁) (h₂ : W.nonsingular x₂ y₂) + (h₁' : W.equation x₁ y₁) (h₂' : W.equation x₂ y₂) + +include h₁ + +lemma XY_ideal_neg_mul : XY_ideal W x₁ (C $ W.neg_Y x₁ y₁) * XY_ideal W x₁ (C y₁) = X_ideal W x₁ := +begin + have Y_rw : + (Y - C (C y₁)) * (Y - C (C (W.neg_Y x₁ y₁))) - C (X - C x₁) + * (C (X ^ 2 + C (x₁ + W.a₂) * X + C (x₁ ^ 2 + W.a₂ * x₁ + W.a₄)) - C (C W.a₁) * Y) + = W.polynomial * 1 := + by linear_combination congr_arg C (congr_arg C ((W.equation_iff _ _).mp h₁.left).symm) + with { normalization_tactic := `[rw [neg_Y, weierstrass_curve.polynomial], C_simp, ring1] }, + simp_rw [XY_ideal, X_class, Y_class, span_pair_mul_span_pair, mul_comm, ← _root_.map_mul, + adjoin_root.mk_eq_mk.mpr ⟨1, Y_rw⟩, _root_.map_mul, span_insert, + ← span_singleton_mul_span_singleton, ← mul_sup, ← span_insert], + convert mul_top _ using 2, + simp_rw [← @set.image_singleton _ _ $ adjoin_root.mk _, ← set.image_insert_eq, ← map_span], + convert map_top (adjoin_root.mk W.polynomial) using 1, + apply congr_arg, + simp_rw [eq_top_iff_one, mem_span_insert', mem_span_singleton'], + cases ((W.nonsingular_iff' _ _).mp h₁).right with hx hy, + { let W_X := W.a₁ * y₁ - (3 * x₁ ^ 2 + 2 * W.a₂ * x₁ + W.a₄), + refine ⟨C (C W_X⁻¹ * -(X + C (2 * x₁ + W.a₂))), C (C $ W_X⁻¹ * W.a₁), 0, C (C $ W_X⁻¹ * -1), _⟩, + rw [← mul_right_inj' $ C_ne_zero.mpr $ C_ne_zero.mpr hx], + simp only [mul_add, ← mul_assoc, ← C_mul, mul_inv_cancel hx], + C_simp, + ring1 }, + { let W_Y := 2 * y₁ + W.a₁ * x₁ + W.a₃, + refine ⟨0, C (C W_Y⁻¹), C (C $ W_Y⁻¹ * -1), 0, _⟩, + rw [neg_Y, ← mul_right_inj' $ C_ne_zero.mpr $ C_ne_zero.mpr hy], + simp only [mul_add, ← mul_assoc, ← C_mul, mul_inv_cancel hy], + C_simp, + ring1 } +end + +private lemma XY_ideal'_mul_inv : + (XY_ideal W x₁ (C y₁) : fractional_ideal W.coordinate_ring⁰ W.function_field) + * (XY_ideal W x₁ (C $ W.neg_Y x₁ y₁) * (X_ideal W x₁)⁻¹) = 1 := +by rw [← mul_assoc, ← fractional_ideal.coe_ideal_mul, mul_comm $ XY_ideal W _ _, + XY_ideal_neg_mul h₁, X_ideal, + fractional_ideal.coe_ideal_span_singleton_mul_inv W.function_field $ X_class_ne_zero W x₁] + +omit h₁ + +include h₁' h₂' + +lemma XY_ideal_mul_XY_ideal (hxy : x₁ = x₂ → y₁ ≠ W.neg_Y x₂ y₂) : + X_ideal W (W.add_X x₁ x₂ $ W.slope x₁ x₂ y₁ y₂) * (XY_ideal W x₁ (C y₁) * XY_ideal W x₂ (C y₂)) + = Y_ideal W (line_polynomial x₁ y₁ $ W.slope x₁ x₂ y₁ y₂) + * XY_ideal W (W.add_X x₁ x₂ $ W.slope x₁ x₂ y₁ y₂) + (C $ W.add_Y x₁ x₂ y₁ $ W.slope x₁ x₂ y₁ y₂) := +begin + have sup_rw : ∀ a b c d : ideal W.coordinate_ring, a ⊔ (b ⊔ (c ⊔ d)) = a ⊔ d ⊔ b ⊔ c := + λ _ _ c _, by rw [← sup_assoc, @sup_comm _ _ c, sup_sup_sup_comm, ← sup_assoc], + rw [XY_ideal_add_eq, X_ideal, mul_comm, W.XY_ideal_eq₁ x₁ y₁ $ W.slope x₁ x₂ y₁ y₂, XY_ideal, + XY_ideal_eq₂ h₁' h₂' hxy, XY_ideal, span_pair_mul_span_pair], + simp_rw [span_insert, sup_rw, sup_mul, span_singleton_mul_span_singleton], + rw [← neg_eq_iff_eq_neg.mpr $ coordinate_ring.C_add_polynomial_slope h₁' h₂' hxy, + span_singleton_neg, coordinate_ring.C_add_polynomial, _root_.map_mul, Y_class], + simp_rw [mul_comm $ X_class W x₁, mul_assoc, ← span_singleton_mul_span_singleton, ← mul_sup], + rw [span_singleton_mul_span_singleton, ← span_insert, + ← span_pair_add_mul_right $ -(X_class W $ W.add_X x₁ x₂ $ W.slope x₁ x₂ y₁ y₂), mul_neg, + ← sub_eq_add_neg, ← sub_mul, ← map_sub, sub_sub_sub_cancel_right, span_insert, + ← span_singleton_mul_span_singleton, ← sup_rw, ← sup_mul, ← sup_mul], + apply congr_arg (_ ∘ _), + convert top_mul _, + simp_rw [X_class, ← @set.image_singleton _ _ $ adjoin_root.mk _, ← map_span, ← ideal.map_sup, + eq_top_iff_one, mem_map_iff_of_surjective _ $ adjoin_root.mk_surjective + W.monic_polynomial, ← span_insert, mem_span_insert', mem_span_singleton'], + by_cases hx : x₁ = x₂, + { rcases ⟨hx, Y_eq_of_Y_ne h₁' h₂' hx (hxy hx)⟩ with ⟨rfl, rfl⟩, + let y := (y₁ - W.neg_Y x₁ y₁) ^ 2, + replace hxy := pow_ne_zero 2 (sub_ne_zero_of_ne $ hxy rfl), + refine + ⟨1 + C (C $ y⁻¹ * 4) * W.polynomial, + ⟨C $ C y⁻¹ * (C 4 * X ^ 2 + C (4 * x₁ + W.b₂) * X + C (4 * x₁ ^ 2 + W.b₂ * x₁ + 2 * W.b₄)), + 0, C (C y⁻¹) * (Y - W.neg_polynomial), _⟩, + by rw [map_add, map_one, _root_.map_mul, adjoin_root.mk_self, mul_zero, add_zero]⟩, + rw [weierstrass_curve.polynomial, neg_polynomial, + ← mul_right_inj' $ C_ne_zero.mpr $ C_ne_zero.mpr hxy], + simp only [mul_add, ← mul_assoc, ← C_mul, mul_inv_cancel hxy], + linear_combination -4 * congr_arg C (congr_arg C $ (W.equation_iff _ _).mp h₁') + with { normalization_tactic := `[rw [b₂, b₄, neg_Y], C_simp, ring1] } }, + { replace hx := sub_ne_zero_of_ne hx, + refine ⟨_, ⟨⟨C $ C (x₁ - x₂)⁻¹, C $ C $ (x₁ - x₂)⁻¹ * -1, 0, _⟩, map_one _⟩⟩, + rw [← mul_right_inj' $ C_ne_zero.mpr $ C_ne_zero.mpr hx], + simp only [← mul_assoc, mul_add, ← C_mul, mul_inv_cancel hx], + C_simp, + ring1 } +end + +omit h₁' h₂' + +/-- The non-zero fractional ideal $\langle X - x, Y - y \rangle$ of $F(W)$ for some $x, y \in F$. -/ +@[simp] noncomputable def XY_ideal' : (fractional_ideal W.coordinate_ring⁰ W.function_field)ˣ := +units.mk_of_mul_eq_one _ _ $ XY_ideal'_mul_inv h₁ + +lemma XY_ideal'_eq : + (XY_ideal' h₁ : fractional_ideal W.coordinate_ring⁰ W.function_field) = XY_ideal W x₁ (C y₁) := +rfl + +local attribute [irreducible] coordinate_ring.comm_ring + +lemma mk_XY_ideal'_mul_mk_XY_ideal'_of_Y_eq : + class_group.mk (XY_ideal' $ nonsingular_neg h₁) * class_group.mk (XY_ideal' h₁) = 1 := +begin + rw [← _root_.map_mul], + exact (class_group.mk_eq_one_of_coe_ideal $ + by exact (fractional_ideal.coe_ideal_mul _ _).symm.trans + (fractional_ideal.coe_ideal_inj.mpr $ XY_ideal_neg_mul h₁)).mpr + ⟨_, X_class_ne_zero W _, rfl⟩ +end + +lemma mk_XY_ideal'_mul_mk_XY_ideal' (hxy : x₁ = x₂ → y₁ ≠ W.neg_Y x₂ y₂) : + class_group.mk (XY_ideal' h₁) * class_group.mk (XY_ideal' h₂) + = class_group.mk (XY_ideal' $ nonsingular_add h₁ h₂ hxy) := +begin + rw [← _root_.map_mul], + exact (class_group.mk_eq_mk_of_coe_ideal (by exact (fractional_ideal.coe_ideal_mul _ _).symm) $ + XY_ideal'_eq _).mpr ⟨_, _, X_class_ne_zero W _, Y_class_ne_zero W _, + XY_ideal_mul_XY_ideal h₁.left h₂.left hxy⟩ +end namespace point +/-- The set function mapping an affine point $(x, y)$ of `W` to the class of the non-zero fractional +ideal $\langle X - x, Y - y \rangle$ of $F(W)$ in the class group of $F[W]$. -/ +@[simp] noncomputable def to_class_fun : W.point → additive (class_group W.coordinate_ring) +| 0 := 0 +| (some h) := additive.of_mul $ class_group.mk $ XY_ideal' h + +/-- The group homomorphism mapping an affine point $(x, y)$ of `W` to the class of the non-zero +fractional ideal $\langle X - x, Y - y \rangle$ of $F(W)$ in the class group of $F[W]$. -/ +@[simps] noncomputable def to_class : W.point →+ additive (class_group W.coordinate_ring) := +{ to_fun := to_class_fun, + map_zero' := rfl, + map_add' := + begin + rintro (_ | @⟨x₁, y₁, h₁⟩) (_ | @⟨x₂, y₂, h₂⟩), + any_goals { simp only [zero_def, to_class_fun, _root_.zero_add, _root_.add_zero] }, + by_cases hx : x₁ = x₂, + { by_cases hy : y₁ = W.neg_Y x₂ y₂, + { substs hx hy, + simpa only [some_add_some_of_Y_eq rfl rfl] + using (mk_XY_ideal'_mul_mk_XY_ideal'_of_Y_eq h₂).symm }, + { simpa only [some_add_some_of_Y_ne hx hy] + using (mk_XY_ideal'_mul_mk_XY_ideal' h₁ h₂ $ λ _, hy).symm } }, + { simpa only [some_add_some_of_X_ne hx] + using (mk_XY_ideal'_mul_mk_XY_ideal' h₁ h₂ $ λ h, (hx h).elim).symm } + end } + +@[simp] lemma to_class_zero : to_class (0 : W.point) = 0 := rfl + +lemma to_class_some : to_class (some h₁) = class_group.mk (XY_ideal' h₁) := rfl + @[simp] lemma add_eq_zero (P Q : W.point) : P + Q = 0 ↔ P = -Q := begin rcases ⟨P, Q⟩ with ⟨_ | @⟨x₁, y₁, _⟩, _ | @⟨x₂, y₂, _⟩⟩, any_goals { refl }, - { rw [zero_def, zero_add, ← neg_eq_iff_eq_neg, neg_zero, eq_comm], }, + { rw [zero_def, zero_add, ← neg_eq_iff_eq_neg, neg_zero, eq_comm] }, { simp only [neg_some], split, { intro h, @@ -528,6 +747,42 @@ end @[simp] lemma neg_add_eq_zero (P Q : W.point) : -P + Q = 0 ↔ P = Q := by rw [add_eq_zero, neg_inj] +lemma to_class_eq_zero (P : W.point) : to_class P = 0 ↔ P = 0 := +⟨begin + intro hP, + rcases P with (_ | @⟨_, _, ⟨h, _⟩⟩), + { refl }, + { rcases (class_group.mk_eq_one_of_coe_ideal $ by refl).mp hP with ⟨p, h0, hp⟩, + apply (p.nat_degree_norm_ne_one _).elim, + rw [← finrank_quotient_span_eq_nat_degree_norm W^.coordinate_ring.basis h0, + ← (quotient_equiv_alg_of_eq F hp).to_linear_equiv.finrank_eq, + (quotient_XY_ideal_equiv W h).to_linear_equiv.finrank_eq, finite_dimensional.finrank_self] } +end, congr_arg to_class⟩ + +lemma to_class_injective : function.injective $ @to_class _ _ W := +begin + rintro (_ | h) _ hP, + all_goals { rw [← neg_add_eq_zero, ← to_class_eq_zero, map_add, ← hP] }, + { exact zero_add 0 }, + { exact mk_XY_ideal'_mul_mk_XY_ideal'_of_Y_eq h } +end + +lemma add_comm (P Q : W.point) : P + Q = Q + P := +to_class_injective $ by simp only [map_add, add_comm] + +lemma add_assoc (P Q R : W.point) : P + Q + R = P + (Q + R) := +to_class_injective $ by simp only [map_add, add_assoc] + +noncomputable instance : add_comm_group W.point := +{ zero := zero, + neg := neg, + add := add, + zero_add := zero_add, + add_zero := add_zero, + add_left_neg := add_left_neg, + add_comm := add_comm, + add_assoc := add_assoc } + end point end group diff --git a/src/algebraic_geometry/elliptic_curve/weierstrass.lean b/src/algebraic_geometry/elliptic_curve/weierstrass.lean index 70aeba4ce34d4..bd0c781db4fb0 100644 --- a/src/algebraic_geometry/elliptic_curve/weierstrass.lean +++ b/src/algebraic_geometry/elliptic_curve/weierstrass.lean @@ -11,7 +11,7 @@ import tactic.linear_combination /-! # Weierstrass equations of elliptic curves -We give a working definition of an elliptic curve as a nonsingular Weierstrass curve given by a +This file defines the structure of an elliptic curve as a nonsingular Weierstrass curve given by a Weierstrass equation, which is mathematically accurate in many cases but also good for computation. ## Mathematical background @@ -257,7 +257,7 @@ section polynomial open polynomial -open_locale polynomial +open_locale polynomial polynomial_polynomial /-- The polynomial $W(X, Y) := Y^2 + a_1XY + a_3Y - (X^3 + a_2X^2 + a_4X + a_6)$ associated to a Weierstrass curve `W` over `R`. For ease of polynomial manipulation, this is represented as a term @@ -440,9 +440,11 @@ abbreviation function_field : Type u := fraction_ring W.coordinate_ring namespace coordinate_ring +open ideal + instance [is_domain R] [normalized_gcd_monoid R] : is_domain W.coordinate_ring := -(ideal.quotient.is_domain_iff_prime _).mpr $ -by simpa only [ideal.span_singleton_prime W.polynomial_ne_zero, ← gcd_monoid.irreducible_iff_prime] +(quotient.is_domain_iff_prime _).mpr $ +by simpa only [span_singleton_prime W.polynomial_ne_zero, ← gcd_monoid.irreducible_iff_prime] using W.irreducible_polynomial instance is_domain_of_field {F : Type u} [field F] (W : weierstrass_curve F) : @@ -466,21 +468,36 @@ adjoin_root.mk_ne_zero_of_nat_degree_lt W.monic_polynomial (X_sub_C_ne_zero y) $ by { rw [nat_degree_polynomial, nat_degree_X_sub_C], norm_num1 } /-- The ideal $\langle X - x \rangle$ of $R[W]$ for some $x \in R$. -/ -@[simp] noncomputable def X_ideal : ideal W.coordinate_ring := ideal.span {X_class W x} +@[simp] noncomputable def X_ideal : ideal W.coordinate_ring := span {X_class W x} /-- The ideal $\langle Y - y(X) \rangle$ of $R[W]$ for some $y(X) \in R[X]$. -/ -@[simp] noncomputable def Y_ideal : ideal W.coordinate_ring := ideal.span {Y_class W y} +@[simp] noncomputable def Y_ideal : ideal W.coordinate_ring := span {Y_class W y} + +/-- The ideal $\langle X - x, Y - y(X) \rangle$ of $R[W]$ for some $x \in R$ and $y(X) \in R[X]$. -/ +@[simp] noncomputable def XY_ideal (x : R) (y : R[X]) : ideal W.coordinate_ring := +span {X_class W x, Y_class W y} /-! ### The coordinate ring as an `R[X]`-algebra -/ -noncomputable instance : algebra R[X] W.coordinate_ring := ideal.quotient.algebra R[X] +noncomputable instance : algebra R[X] W.coordinate_ring := quotient.algebra R[X] -noncomputable instance algebra' : algebra R W.coordinate_ring := ideal.quotient.algebra R +noncomputable instance algebra' : algebra R W.coordinate_ring := quotient.algebra R -instance : is_scalar_tower R R[X] W.coordinate_ring := ideal.quotient.is_scalar_tower R R[X] _ +instance : is_scalar_tower R R[X] W.coordinate_ring := quotient.is_scalar_tower R R[X] _ instance [subsingleton R] : subsingleton W.coordinate_ring := module.subsingleton R[X] _ +/-- The $R$-algebra isomorphism from $R[W] / \langle X - x, Y - y(X) \rangle$ to $R$ obtained by +evaluation at $y(X)$ and at $x$ provided that $W(x, y(x)) = 0$. -/ +noncomputable def quotient_XY_ideal_equiv {x : R} {y : R[X]} + (h : (W.polynomial.eval y).eval x = 0) : (W.coordinate_ring ⧸ XY_ideal W x y) ≃ₐ[R] R := +(quotient_equiv_alg_of_eq R $ + by simpa only [XY_ideal, X_class, Y_class, ← set.image_pair, ← map_span]).trans $ + (double_quot.quot_quot_equiv_quot_of_leₐ R $ (span_singleton_le_iff_mem _).mpr $ + mem_span_C_X_sub_C_X_sub_C_iff_eval_eval_eq_zero.mpr h).trans $ + ((quotient_span_C_X_sub_C_alg_equiv (X - C x) y).restrict_scalars R).trans $ + quotient_span_X_sub_C_alg_equiv x + /-- The basis $\{1, Y\}$ for the coordinate ring $R[W]$ over the polynomial ring $R[X]$. Given a Weierstrass curve `W`, write `W^.coordinate_ring.basis` for this basis. -/ @@ -533,7 +550,7 @@ variable (W) lemma smul_basis_mul_C (p q : R[X]) : (p • 1 + q • adjoin_root.mk W.polynomial Y) * adjoin_root.mk W.polynomial (C y) = ((p * y) • 1 + (q * y) • adjoin_root.mk W.polynomial Y) := -by { simp only [smul, map_mul], ring1 } +by { simp only [smul, _root_.map_mul], ring1 } lemma smul_basis_mul_Y (p q : R[X]) : (p • 1 + q • adjoin_root.mk W.polynomial Y) * adjoin_root.mk W.polynomial Y @@ -543,7 +560,7 @@ begin have Y_sq : adjoin_root.mk W.polynomial Y ^ 2 = adjoin_root.mk W.polynomial (C (X ^ 3 + C W.a₂ * X ^ 2 + C W.a₄ * X + C W.a₆) - C (C W.a₁ * X + C W.a₃) * Y) := adjoin_root.mk_eq_mk.mpr ⟨1, by { simp only [weierstrass_curve.polynomial], ring1 }⟩, - simp only [smul, add_mul, mul_assoc, ← sq, Y_sq, map_sub, map_mul], + simp only [smul, add_mul, mul_assoc, ← sq, Y_sq, map_sub, _root_.map_mul], ring1 end @@ -564,7 +581,7 @@ end lemma coe_norm_smul_basis (p q : R[X]) : ↑(algebra.norm R[X] $ p • 1 + q • adjoin_root.mk W.polynomial Y) = adjoin_root.mk W.polynomial - ((C p + C q * X) * (C p + C q * (-X - C (C W.a₁ * X + C W.a₃)))) := + ((C p + C q * X) * (C p + C q * (-Y - C (C W.a₁ * X + C W.a₃)))) := adjoin_root.mk_eq_mk.mpr ⟨C q ^ 2, by { rw [norm_smul_basis, weierstrass_curve.polynomial], C_simp, ring1 }⟩ From fbbd626e4e8fe4d1ce64903e26276ec2873f2976 Mon Sep 17 00:00:00 2001 From: Thomas Browning Date: Tue, 13 Jun 2023 07:06:59 +0000 Subject: [PATCH 1179/1291] chore(field_theory/abel_ruffini): Rename variables (#19180) Rename variables for consistency. --- src/field_theory/abel_ruffini.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/field_theory/abel_ruffini.lean b/src/field_theory/abel_ruffini.lean index 0909b224c515d..3f8412f818bab 100644 --- a/src/field_theory/abel_ruffini.lean +++ b/src/field_theory/abel_ruffini.lean @@ -209,8 +209,8 @@ variables (F) /-- Inductive definition of solvable by radicals -/ inductive is_solvable_by_rad : E → Prop -| base (a : F) : is_solvable_by_rad (algebra_map F E a) -| add (a b : E) : is_solvable_by_rad a → is_solvable_by_rad b → is_solvable_by_rad (a + b) +| base (α : F) : is_solvable_by_rad (algebra_map F E α) +| add (α β : E) : is_solvable_by_rad α → is_solvable_by_rad β → is_solvable_by_rad (α + β) | neg (α : E) : is_solvable_by_rad α → is_solvable_by_rad (-α) | mul (α β : E) : is_solvable_by_rad α → is_solvable_by_rad β → is_solvable_by_rad (α * β) | inv (α : E) : is_solvable_by_rad α → is_solvable_by_rad α⁻¹ From 8efcf8022aac8e01df8d302dcebdbc25d6a886c8 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 13 Jun 2023 07:07:01 +0000 Subject: [PATCH 1180/1291] refactor(data/polynomial/ring_division): remove `open_locale classical` (#19182) This makes the lemmas strictly more general. --- src/data/polynomial/ring_division.lean | 70 ++++++++++++++-------- src/linear_algebra/eigenspace/minpoly.lean | 1 + 2 files changed, 47 insertions(+), 24 deletions(-) diff --git a/src/data/polynomial/ring_division.lean b/src/data/polynomial/ring_division.lean index d833e027586a3..a9c8a85e0f496 100644 --- a/src/data/polynomial/ring_division.lean +++ b/src/data/polynomial/ring_division.lean @@ -33,7 +33,7 @@ This file starts looking like the ring theory of $ R[X] $ -/ noncomputable theory -open_locale classical polynomial +open_locale polynomial open finset @@ -141,6 +141,7 @@ end @[simp] lemma nat_degree_pow (p : R[X]) (n : ℕ) : nat_degree (p ^ n) = n * nat_degree p := +by classical; exact if hp0 : p = 0 then if hn0 : n = 0 then by simp [hp0, hn0] else by rw [hp0, zero_pow (nat.pos_of_ne_zero hn0)]; simp @@ -148,6 +149,7 @@ else nat_degree_pow' (by rw [← leading_coeff_pow, ne.def, leading_coeff_eq_zero]; exact pow_ne_zero _ hp0) lemma degree_le_mul_left (p : R[X]) (hq : q ≠ 0) : degree p ≤ degree (p * q) := +by classical; exact if hp : p = 0 then by simp only [hp, zero_mul, le_refl] else by rw [degree_mul, degree_eq_nat_degree hp, degree_eq_nat_degree hq]; @@ -338,6 +340,7 @@ variable [comm_ring R] lemma le_root_multiplicity_iff {p : R[X]} (p0 : p ≠ 0) {a : R} {n : ℕ} : n ≤ root_multiplicity a p ↔ (X - C a) ^ n ∣ p := begin + classical, simp_rw [root_multiplicity, dif_neg p0, nat.le_find_iff, not_not], refine ⟨λ h, _, λ h m hm, (pow_dvd_pow _ hm).trans h⟩, cases n, { rw pow_zero, apply one_dvd }, { exact h n n.lt_succ_self }, @@ -403,6 +406,7 @@ end lemma root_multiplicity_mul {p q : R[X]} {x : R} (hpq : p * q ≠ 0) : root_multiplicity x (p * q) = root_multiplicity x p + root_multiplicity x q := begin + classical, have hp : p ≠ 0 := left_ne_zero_of_mul hpq, have hq : q ≠ 0 := right_ne_zero_of_mul hpq, rw [root_multiplicity_eq_multiplicity (p * q), dif_neg hpq, @@ -413,10 +417,10 @@ end lemma root_multiplicity_X_sub_C_self {x : R} : root_multiplicity x (X - C x) = 1 := -by rw [root_multiplicity_eq_multiplicity, dif_neg (X_sub_C_ne_zero x), +by classical; rw [root_multiplicity_eq_multiplicity, dif_neg (X_sub_C_ne_zero x), multiplicity.get_multiplicity_self] -lemma root_multiplicity_X_sub_C {x y : R} : +lemma root_multiplicity_X_sub_C {x y : R} [decidable_eq R] : root_multiplicity x (X - C y) = if x = y then 1 else 0 := begin split_ifs with hxy, @@ -436,7 +440,7 @@ begin simp only [root_multiplicity_mul hzero, root_multiplicity_X_sub_C_self, hn, nat.one_add] end -lemma exists_multiset_roots : ∀ {p : R[X]} (hp : p ≠ 0), +lemma exists_multiset_roots [decidable_eq R] : ∀ {p : R[X]} (hp : p ≠ 0), ∃ s : multiset R, (s.card : with_bot ℕ) ≤ degree p ∧ ∀ a, s.count a = root_multiplicity a p | p := λ hp, by haveI := classical.prop_decidable (∃ x, is_root p x); exact if h : ∃ x, is_root p x @@ -479,13 +483,24 @@ using_well_founded {dec_tac := tactic.assumption} /-- `roots p` noncomputably gives a multiset containing all the roots of `p`, including their multiplicities. -/ noncomputable def roots (p : R[X]) : multiset R := -if h : p = 0 then ∅ else classical.some (exists_multiset_roots h) +by haveI := classical.dec_eq R; haveI := classical.dec (p = 0); exact + if h : p = 0 then ∅ else classical.some (exists_multiset_roots h) + +lemma roots_def [decidable_eq R] (p : R[X]) [decidable (p = 0)] : + p.roots = if h : p = 0 then ∅ else classical.some (exists_multiset_roots h) := +begin + unfreezingI + { obtain rfl := subsingleton.elim ‹_› (classical.dec_eq R), + obtain rfl := subsingleton.elim ‹_› (classical.dec (p = 0)),}, + refl, +end @[simp] lemma roots_zero : (0 : R[X]).roots = 0 := -dif_pos rfl +by apply dif_pos rfl lemma card_roots (hp0 : p ≠ 0) : ((roots p).card : with_bot ℕ) ≤ degree p := begin + classical, unfold roots, rw dif_neg hp0, exact (classical.some_spec (exists_multiset_roots hp0)).1 @@ -509,16 +524,17 @@ lemma card_roots_sub_C' {p : R[X]} {a : R} (hp0 : 0 < degree p) : with_bot.coe_le_coe.1 (le_trans (card_roots_sub_C hp0) (le_of_eq $ degree_eq_nat_degree (λ h, by simp [*, lt_irrefl] at *))) -@[simp] lemma count_roots (p : R[X]) : p.roots.count a = root_multiplicity a p := +@[simp] lemma count_roots [decidable_eq R] (p : R[X]) : p.roots.count a = root_multiplicity a p := begin + classical, by_cases hp : p = 0, { simp [hp], }, - rw [roots, dif_neg hp], - exact (classical.some_spec (exists_multiset_roots hp)).2 a + rw [roots_def, dif_neg hp], + exact (classical.some_spec (exists_multiset_roots hp)).2 a, end @[simp] lemma mem_roots' : a ∈ p.roots ↔ p ≠ 0 ∧ is_root p a := -by rw [← count_pos, count_roots p, root_multiplicity_pos'] +by classical; rw [← count_pos, count_roots p, root_multiplicity_pos'] lemma mem_roots (hp : p ≠ 0) : a ∈ p.roots ↔ is_root p a := mem_roots'.trans $ and_iff_right hp @@ -531,7 +547,7 @@ theorem card_le_degree_of_subset_roots {p : R[X]} {Z : finset R} (h : Z.val ⊆ (multiset.card_le_of_le (finset.val_le_iff_val_subset.2 h)).trans (polynomial.card_roots' p) lemma finite_set_of_is_root {p : R[X]} (hp : p ≠ 0) : set.finite {x | is_root p x} := -by simpa only [← finset.set_of_mem, mem_to_finset, mem_roots hp] +by classical; simpa only [← finset.set_of_mem, mem_to_finset, mem_roots hp] using p.roots.to_finset.finite_to_set lemma eq_zero_of_infinite_is_root (p : R[X]) (h : set.infinite {x | is_root p x}) : p = 0 := @@ -553,9 +569,9 @@ begin end lemma roots_mul {p q : R[X]} (hpq : p * q ≠ 0) : (p * q).roots = p.roots + q.roots := -multiset.ext.mpr $ λ r, +by classical; exact (multiset.ext.mpr $ λ r, by rw [count_add, count_roots, count_roots, - count_roots, root_multiplicity_mul hpq] + count_roots, root_multiplicity_mul hpq]) lemma roots.le_of_dvd (h : q ≠ 0) : p ∣ q → roots p ≤ roots q := begin @@ -573,6 +589,7 @@ mem_roots_sub_C'.trans $ and_iff_right $ λ hp, hp0.not_le $ hp.symm ▸ degree_ @[simp] lemma roots_X_sub_C (r : R) : roots (X - C r) = {r} := begin + classical, ext s, rw [count_roots, root_multiplicity_X_sub_C, count_singleton], end @@ -580,7 +597,7 @@ end @[simp] lemma roots_X : roots (X : R[X]) = {0} := by rw [← roots_X_sub_C, C_0, sub_zero] @[simp] lemma roots_C (x : R) : (C x).roots = 0 := -if H : x = 0 then by rw [H, C_0, roots_zero] else multiset.ext.mpr $ λ r, +by classical; exact if H : x = 0 then by rw [H, C_0, roots_zero] else multiset.ext.mpr $ λ r, by rw [count_roots, count_zero, root_multiplicity_eq_zero (not_is_root_C _ _ H)] @[simp] lemma roots_one : (1 : R[X]).roots = ∅ := @@ -676,6 +693,7 @@ by simp only [empty_eq_zero, pow_zero, nth_roots, ← C_1, ← C_sub, roots_C] lemma card_nth_roots (n : ℕ) (a : R) : (nth_roots n a).card ≤ n := +by classical; exactI if hn : n = 0 then if h : (X : R[X]) ^ n - C a = 0 then by simp only [nat.zero_le, nth_roots, roots, h, dif_pos rfl, empty_eq_zero, multiset.card_zero] @@ -692,7 +710,7 @@ by simp_rw [is_square_iff_exists_sq, eq_zero_iff_forall_not_mem, /-- The multiset `nth_roots ↑n (1 : R)` as a finset. -/ def nth_roots_finset (n : ℕ) (R : Type*) [comm_ring R] [is_domain R] : finset R := -multiset.to_finset (nth_roots n (1 : R)) +by haveI := classical.dec_eq R; exact multiset.to_finset (nth_roots n (1 : R)) @[simp] lemma mem_nth_roots_finset {n : ℕ} (h : 0 < n) {x : R} : x ∈ nth_roots_finset n R ↔ x ^ (n : ℕ) = 1 := @@ -758,15 +776,15 @@ variables [comm_ring T] If you have a non-separable polynomial, use `polynomial.roots` for the multiset where multiple roots have the appropriate multiplicity. -/ def root_set (p : T[X]) (S) [comm_ring S] [is_domain S] [algebra T S] : set S := -(p.map (algebra_map T S)).roots.to_finset +by haveI := classical.dec_eq S; exact (p.map (algebra_map T S)).roots.to_finset -lemma root_set_def (p : T[X]) (S) [comm_ring S] [is_domain S] [algebra T S] : +lemma root_set_def (p : T[X]) (S) [comm_ring S] [is_domain S] [algebra T S] [decidable_eq S] : p.root_set S = (p.map (algebra_map T S)).roots.to_finset := -rfl +by convert rfl @[simp] lemma root_set_C [comm_ring S] [is_domain S] [algebra T S] (a : T) : (C a).root_set S = ∅ := -by rw [root_set_def, map_C, roots_C, multiset.to_finset_zero, finset.coe_empty] +by classical; rw [root_set_def, map_C, roots_C, multiset.to_finset_zero, finset.coe_empty] @[simp] lemma root_set_zero (S) [comm_ring S] [is_domain S] [algebra T S] : (0 : T[X]).root_set S = ∅ := @@ -782,7 +800,7 @@ set.to_finite _ /-- The set of roots of all polynomials of bounded degree and having coefficients in a finite set is finite. -/ -lemma bUnion_roots_finite {R S : Type*} [semiring R] [comm_ring S] [is_domain S] +lemma bUnion_roots_finite {R S : Type*} [semiring R] [comm_ring S] [is_domain S] [decidable_eq S] (m : R →+* S) (d : ℕ) {U : set R} (h : U.finite) : (⋃ (f : R[X]) (hf : f.nat_degree ≤ d ∧ ∀ i, (f.coeff i) ∈ U), ((f.map m).roots.to_finset : set S)).finite := @@ -920,7 +938,7 @@ theorem pairwise_coprime_X_sub_C {K} [field K] {I : Type v} {s : I → K} lemma monic_prod_multiset_X_sub_C : monic (p.roots.map (λ a, X - C a)).prod := monic_multiset_prod_of_monic _ _ (λ a _, monic_X_sub_C a) -lemma prod_multiset_root_eq_finset_root : +lemma prod_multiset_root_eq_finset_root [decidable_eq R] : (p.roots.map (λ a, X - C a)).prod = p.roots.to_finset.prod (λ a, (X - C a) ^ root_multiplicity a p) := by simp only [count_roots, finset.prod_multiset_map_count] @@ -928,6 +946,7 @@ by simp only [count_roots, finset.prod_multiset_map_count] /-- The product `∏ (X - a)` for `a` inside the multiset `p.roots` divides `p`. -/ lemma prod_multiset_X_sub_C_dvd (p : R[X]) : (p.roots.map (λ a, X - C a)).prod ∣ p := begin + classical, rw ← map_dvd_map _ (is_fraction_ring.injective R $ fraction_ring R) monic_prod_multiset_X_sub_C, rw [prod_multiset_root_eq_finset_root, polynomial.map_prod], refine finset.prod_dvd_of_coprime (λ a _ b _ h, _) (λ a _, _), @@ -939,6 +958,7 @@ end /-- A Galois connection. -/ lemma _root_.multiset.prod_X_sub_C_dvd_iff_le_roots {p : R[X]} (hp : p ≠ 0) (s : multiset R) : (s.map (λ a, X - C a)).prod ∣ p ↔ s ≤ p.roots := +by classical; exact ⟨λ h, multiset.le_iff_count.2 $ λ r, begin rw [count_roots, le_root_multiplicity_iff hp, ← multiset.prod_replicate, ← multiset.map_replicate (λ a, X - C a), ← multiset.filter_eq], @@ -999,7 +1019,8 @@ begin apply pow_root_multiplicity_dvd, end -lemma count_map_roots [is_domain A] {p : A[X]} {f : A →+* B} (hmap : map f p ≠ 0) (b : B) : +lemma count_map_roots [is_domain A] [decidable_eq B] {p : A[X]} {f : A →+* B} (hmap : map f p ≠ 0) + (b : B) : (p.roots.map f).count b ≤ root_multiplicity b (p.map f) := begin rw [le_root_multiplicity_iff hmap, ← multiset.prod_replicate, @@ -1012,7 +1033,7 @@ begin simp only [function.comp_app, polynomial.map_sub, map_X, map_C], end -lemma count_map_roots_of_injective [is_domain A] (p : A[X]) {f : A →+* B} +lemma count_map_roots_of_injective [is_domain A] [decidable_eq B] (p : A[X]) {f : A →+* B} (hf : function.injective f) (b : B) : (p.roots.map f).count b ≤ root_multiplicity b (p.map f) := begin @@ -1024,7 +1045,8 @@ end lemma map_roots_le [is_domain A] [is_domain B] {p : A[X]} {f : A →+* B} (h : p.map f ≠ 0) : p.roots.map f ≤ (p.map f).roots := -multiset.le_iff_count.2 $ λ b, by { rw count_roots, apply count_map_roots h } +by classical; exact + (multiset.le_iff_count.2 $ λ b, by { rw count_roots, apply count_map_roots h }) lemma map_roots_le_of_injective [is_domain A] [is_domain B] (p : A[X]) {f : A →+* B} (hf : function.injective f) : diff --git a/src/linear_algebra/eigenspace/minpoly.lean b/src/linear_algebra/eigenspace/minpoly.lean index f371c2eadd8ab..1b4dac34fd00f 100644 --- a/src/linear_algebra/eigenspace/minpoly.lean +++ b/src/linear_algebra/eigenspace/minpoly.lean @@ -108,6 +108,7 @@ begin ext μ, have : (μ ∈ {μ : K | f.eigenspace μ = ⊥ → false}) ↔ ¬f.eigenspace μ = ⊥ := by tauto, convert rfl.mpr this, + classical, simp [polynomial.root_set_def, polynomial.mem_roots h, ← has_eigenvalue_iff_is_root, has_eigenvalue] end From d07a9c875ed7139abfde6a333b2be205c5bd404e Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Wed, 14 Jun 2023 04:48:05 +0000 Subject: [PATCH 1181/1291] chore(*): add mathlib4 synchronization comments (#19185) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.category.Group.subobject` * `analysis.analytic.composition` * `analysis.convex.intrinsic` * `analysis.inner_product_space.two_dim` * `analysis.special_functions.japanese_bracket` * `category_theory.differential_object` * `geometry.euclidean.angle.oriented.basic` * `measure_theory.function.conditional_expectation.ae_measurable` * `order.category.BddOrd` * `order.category.Semilat` * `ring_theory.polynomial.hermite.basic` * `ring_theory.polynomial.hermite.gaussian` * `ring_theory.ring_hom.integral` * `ring_theory.ring_hom_properties` --- src/algebra/category/Group/subobject.lean | 3 +++ src/analysis/analytic/composition.lean | 3 +++ src/analysis/convex/intrinsic.lean | 3 +++ src/analysis/inner_product_space/two_dim.lean | 3 +++ src/analysis/special_functions/japanese_bracket.lean | 3 +++ src/category_theory/differential_object.lean | 3 +++ src/geometry/euclidean/angle/oriented/basic.lean | 3 +++ .../function/conditional_expectation/ae_measurable.lean | 3 +++ src/order/category/BddOrd.lean | 3 +++ src/order/category/Semilat.lean | 3 +++ src/ring_theory/polynomial/hermite/basic.lean | 3 +++ src/ring_theory/polynomial/hermite/gaussian.lean | 3 +++ src/ring_theory/ring_hom/integral.lean | 3 +++ src/ring_theory/ring_hom_properties.lean | 3 +++ 14 files changed, 42 insertions(+) diff --git a/src/algebra/category/Group/subobject.lean b/src/algebra/category/Group/subobject.lean index 07506cf2a8e19..47fca85c2d57a 100644 --- a/src/algebra/category/Group/subobject.lean +++ b/src/algebra/category/Group/subobject.lean @@ -8,6 +8,9 @@ import algebra.category.Module.subobject /-! # The category of abelian groups is well-powered + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ open category_theory diff --git a/src/analysis/analytic/composition.lean b/src/analysis/analytic/composition.lean index 72c45ccf16d31..83396029efcdc 100644 --- a/src/analysis/analytic/composition.lean +++ b/src/analysis/analytic/composition.lean @@ -9,6 +9,9 @@ import combinatorics.composition /-! # Composition of analytic functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove that the composition of analytic functions is analytic. The argument is the following. Assume `g z = ∑' qₙ (z, ..., z)` and `f y = ∑' pₖ (y, ..., y)`. Then diff --git a/src/analysis/convex/intrinsic.lean b/src/analysis/convex/intrinsic.lean index 5d36a26c91ffe..f3e91c330fe29 100644 --- a/src/analysis/convex/intrinsic.lean +++ b/src/analysis/convex/intrinsic.lean @@ -8,6 +8,9 @@ import analysis.normed_space.add_torsor_bases /-! # Intrinsic frontier and interior +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the intrinsic frontier, interior and closure of a set in a normed additive torsor. These are also known as relative frontier, interior, closure. diff --git a/src/analysis/inner_product_space/two_dim.lean b/src/analysis/inner_product_space/two_dim.lean index 62fa50cecc64e..f353988543fa0 100644 --- a/src/analysis/inner_product_space/two_dim.lean +++ b/src/analysis/inner_product_space/two_dim.lean @@ -11,6 +11,9 @@ import tactic.linear_combination /-! # Oriented two-dimensional real inner product spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines constructions specific to the geometry of an oriented two-dimensional real inner product space `E`. diff --git a/src/analysis/special_functions/japanese_bracket.lean b/src/analysis/special_functions/japanese_bracket.lean index 94cf135552127..6108b5d232139 100644 --- a/src/analysis/special_functions/japanese_bracket.lean +++ b/src/analysis/special_functions/japanese_bracket.lean @@ -9,6 +9,9 @@ import measure_theory.integral.layercake /-! # Japanese Bracket +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we show that Japanese bracket $(1 + \|x\|^2)^{1/2}$ can be estimated from above and below by $1 + \|x\|$. The functions $(1 + \|x\|^2)^{-r/2}$ and $(1 + |x|)^{-r}$ are integrable provided that `r` is larger diff --git a/src/category_theory/differential_object.lean b/src/category_theory/differential_object.lean index 345d26ac9dd25..38865e33667c7 100644 --- a/src/category_theory/differential_object.lean +++ b/src/category_theory/differential_object.lean @@ -10,6 +10,9 @@ import category_theory.concrete_category.basic /-! # Differential objects in a category. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A differential object in a category with zero morphisms and a shift is an object `X` equipped with a morphism `d : X ⟶ X⟦1⟧`, such that `d^2 = 0`. diff --git a/src/geometry/euclidean/angle/oriented/basic.lean b/src/geometry/euclidean/angle/oriented/basic.lean index 34957f68fe9b2..b5de68815a179 100644 --- a/src/geometry/euclidean/angle/oriented/basic.lean +++ b/src/geometry/euclidean/angle/oriented/basic.lean @@ -9,6 +9,9 @@ import geometry.euclidean.angle.unoriented.basic /-! # Oriented angles. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines oriented angles in real inner product spaces. ## Main definitions diff --git a/src/measure_theory/function/conditional_expectation/ae_measurable.lean b/src/measure_theory/function/conditional_expectation/ae_measurable.lean index ed446146e1a82..2cc7b7b1d1eb7 100644 --- a/src/measure_theory/function/conditional_expectation/ae_measurable.lean +++ b/src/measure_theory/function/conditional_expectation/ae_measurable.lean @@ -8,6 +8,9 @@ import measure_theory.function.strongly_measurable.lp /-! # Functions a.e. measurable with respect to a sub-σ-algebra +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A function `f` verifies `ae_strongly_measurable' m f μ` if it is `μ`-a.e. equal to an `m`-strongly measurable function. This is similar to `ae_strongly_measurable`, but the `measurable_space` structures used for the measurability statement and for the measure are diff --git a/src/order/category/BddOrd.lean b/src/order/category/BddOrd.lean index 0be06702f6a2f..b76686272a796 100644 --- a/src/order/category/BddOrd.lean +++ b/src/order/category/BddOrd.lean @@ -10,6 +10,9 @@ import order.hom.bounded /-! # The category of bounded orders +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This defines `BddOrd`, the category of bounded orders. -/ diff --git a/src/order/category/Semilat.lean b/src/order/category/Semilat.lean index 44b058fcc2e7f..49a62906cbed4 100644 --- a/src/order/category/Semilat.lean +++ b/src/order/category/Semilat.lean @@ -9,6 +9,9 @@ import order.hom.lattice /-! # The categories of semilattices +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This defines `SemilatSup` and `SemilatInf`, the categories of sup-semilattices with a bottom element and inf-semilattices with a top element. diff --git a/src/ring_theory/polynomial/hermite/basic.lean b/src/ring_theory/polynomial/hermite/basic.lean index 5df004058baac..8c324c3e962c8 100644 --- a/src/ring_theory/polynomial/hermite/basic.lean +++ b/src/ring_theory/polynomial/hermite/basic.lean @@ -10,6 +10,9 @@ import data.nat.factorial.double_factorial /-! # Hermite polynomials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines `polynomial.hermite n`, the nth probabilist's Hermite polynomial. ## Main definitions diff --git a/src/ring_theory/polynomial/hermite/gaussian.lean b/src/ring_theory/polynomial/hermite/gaussian.lean index 261e7de7865eb..307939d32d9fa 100644 --- a/src/ring_theory/polynomial/hermite/gaussian.lean +++ b/src/ring_theory/polynomial/hermite/gaussian.lean @@ -12,6 +12,9 @@ import analysis.special_functions.exp_deriv /-! # Hermite polynomials and Gaussians +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file shows that the Hermite polynomial `hermite n` is (up to sign) the polynomial factor occurring in the `n`th derivative of a gaussian. diff --git a/src/ring_theory/ring_hom/integral.lean b/src/ring_theory/ring_hom/integral.lean index 22318a63766fa..1d96571c23f5f 100644 --- a/src/ring_theory/ring_hom/integral.lean +++ b/src/ring_theory/ring_hom/integral.lean @@ -10,6 +10,9 @@ import ring_theory.integral_closure # The meta properties of integral ring homomorphisms. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + -/ namespace ring_hom diff --git a/src/ring_theory/ring_hom_properties.lean b/src/ring_theory/ring_hom_properties.lean index 321ccee362f0b..e83e147d9c865 100644 --- a/src/ring_theory/ring_hom_properties.lean +++ b/src/ring_theory/ring_hom_properties.lean @@ -12,6 +12,9 @@ import ring_theory.is_tensor_product /-! # Properties of ring homomorphisms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We provide the basic framework for talking about properties of ring homomorphisms. The following meta-properties of predicates on ring homomorphisms are defined From c3216069e5f9369e6be586ccbfcde2592b3cec92 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 14 Jun 2023 09:28:31 +0000 Subject: [PATCH 1182/1291] chore(linear_algebra/eigenspace/minpoly): remove a silly use of `tauto` (#19183) `tauto` is not `refl`. --- src/linear_algebra/eigenspace/minpoly.lean | 8 +++----- 1 file changed, 3 insertions(+), 5 deletions(-) diff --git a/src/linear_algebra/eigenspace/minpoly.lean b/src/linear_algebra/eigenspace/minpoly.lean index 1b4dac34fd00f..ada135c892033 100644 --- a/src/linear_algebra/eigenspace/minpoly.lean +++ b/src/linear_algebra/eigenspace/minpoly.lean @@ -101,16 +101,14 @@ theorem has_eigenvalue_iff_is_root : /-- An endomorphism of a finite-dimensional vector space has finitely many eigenvalues. -/ noncomputable instance (f : End K V) : fintype f.eigenvalues := -set.finite.fintype +set.finite.fintype $ show {μ | eigenspace f μ ≠ ⊥}.finite, begin have h : minpoly K f ≠ 0 := minpoly.ne_zero f.is_integral, - convert (minpoly K f).root_set_finite K, + convert (minpoly K f).root_set_finite K using 1, ext μ, - have : (μ ∈ {μ : K | f.eigenspace μ = ⊥ → false}) ↔ ¬f.eigenspace μ = ⊥ := by tauto, - convert rfl.mpr this, classical, simp [polynomial.root_set_def, polynomial.mem_roots h, ← has_eigenvalue_iff_is_root, - has_eigenvalue] + has_eigenvalue], end end End From 1ac8d4304efba9d03fa720d06516fac845aa5353 Mon Sep 17 00:00:00 2001 From: Eric Rodriguez Date: Wed, 14 Jun 2023 10:32:56 +0000 Subject: [PATCH 1183/1291] chore(*): fix `@[to_additive, simp]` to the correct order (#19169) Whilst making some files, I noticed that there is some lemmas that have the wrong order for `to_additive` and `simp`. --- src/algebra/hom/equiv/basic.lean | 4 ++-- src/group_theory/subsemigroup/center.lean | 2 +- src/topology/algebra/monoid.lean | 4 ++-- 3 files changed, 5 insertions(+), 5 deletions(-) diff --git a/src/algebra/hom/equiv/basic.lean b/src/algebra/hom/equiv/basic.lean index 423615c762cc7..6a616f42e7923 100644 --- a/src/algebra/hom/equiv/basic.lean +++ b/src/algebra/hom/equiv/basic.lean @@ -309,10 +309,10 @@ fun_like.ext _ _ e.apply_symm_apply theorem self_trans_symm (e : M ≃* N) : e.trans e.symm = refl M := fun_like.ext _ _ e.symm_apply_apply -@[to_additive, simp] lemma coe_monoid_hom_refl {M} [mul_one_class M] : +@[simp, to_additive] lemma coe_monoid_hom_refl {M} [mul_one_class M] : (refl M : M →* M) = monoid_hom.id M := rfl -@[to_additive, simp] lemma coe_monoid_hom_trans {M N P} +@[simp, to_additive] lemma coe_monoid_hom_trans {M N P} [mul_one_class M] [mul_one_class N] [mul_one_class P] (e₁ : M ≃* N) (e₂ : N ≃* P) : (e₁.trans e₂ : M →* P) = (e₂ : N →* P).comp ↑e₁ := rfl diff --git a/src/group_theory/subsemigroup/center.lean b/src/group_theory/subsemigroup/center.lean index 0ab9e2901c059..46cd73116a996 100644 --- a/src/group_theory/subsemigroup/center.lean +++ b/src/group_theory/subsemigroup/center.lean @@ -151,7 +151,7 @@ end section variables (M) [comm_semigroup M] -@[to_additive, simp] lemma center_eq_top : center M = ⊤ := +@[simp, to_additive] lemma center_eq_top : center M = ⊤ := set_like.coe_injective (set.center_eq_univ M) end diff --git a/src/topology/algebra/monoid.lean b/src/topology/algebra/monoid.lean index 5dbee1af0e363..305780ebd3702 100644 --- a/src/topology/algebra/monoid.lean +++ b/src/topology/algebra/monoid.lean @@ -698,14 +698,14 @@ variables [has_mul X] [has_continuous_mul X] @[to_additive "The continuous map `λ y, y + x"] protected def mul_right (x : X) : C(X, X) := mk _ (continuous_mul_right x) -@[to_additive, simp] +@[simp, to_additive] lemma coe_mul_right (x : X) : ⇑(continuous_map.mul_right x) = λ y, y * x := rfl /-- The continuous map `λ y, x * y` -/ @[to_additive "The continuous map `λ y, x + y"] protected def mul_left (x : X) : C(X, X) := mk _ (continuous_mul_left x) -@[to_additive, simp] +@[simp, to_additive] lemma coe_mul_left (x : X) : ⇑(continuous_map.mul_left x) = λ y, x * y := rfl end continuous_map From f9ec187127cc5b381dfcf5f4a22dacca4c20b63d Mon Sep 17 00:00:00 2001 From: Floris van Doorn Date: Wed, 14 Jun 2023 13:43:53 +0000 Subject: [PATCH 1184/1291] feat(geometry/manifold/vector_bundle/smooth_section): define smooth sections (#19064) * Define the module of smooth sections of a smooth vector bundle over a smooth manifold. * Co-authored-by: Heather Macbeth [25316162+hrmacbeth@users.noreply.github.com](mailto:25316162+hrmacbeth@users.noreply.github.com) Co-authored-by: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> --- src/geometry/manifold/algebra/lie_group.lean | 84 ++++++- .../manifold/vector_bundle/basic.lean | 6 + .../vector_bundle/smooth_section.lean | 210 ++++++++++++++++++ 3 files changed, 291 insertions(+), 9 deletions(-) create mode 100644 src/geometry/manifold/vector_bundle/smooth_section.lean diff --git a/src/geometry/manifold/algebra/lie_group.lean b/src/geometry/manifold/algebra/lie_group.lean index 7866fc867b9a0..32933bde56c91 100644 --- a/src/geometry/manifold/algebra/lie_group.lean +++ b/src/geometry/manifold/algebra/lie_group.lean @@ -74,7 +74,7 @@ variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] {M : Type*} [topological_space M] [charted_space H' M] {E'' : Type*} [normed_add_comm_group E''] [normed_space 𝕜 E''] {H'' : Type*} [topological_space H''] {I'' : model_with_corners 𝕜 E'' H''} -{M' : Type*} [topological_space M'] [charted_space H'' M'] +{M' : Type*} [topological_space M'] [charted_space H'' M'] {n : ℕ∞} section @@ -96,24 +96,90 @@ lemma topological_group_of_lie_group : topological_group G := end @[to_additive] -lemma smooth.inv {f : M → G} - (hf : smooth I' I f) : smooth I' I (λx, (f x)⁻¹) := -(smooth_inv I).comp hf +lemma cont_mdiff_within_at.inv {f : M → G} {s : set M} {x₀ : M} + (hf : cont_mdiff_within_at I' I n f s x₀) : cont_mdiff_within_at I' I n (λx, (f x)⁻¹) s x₀ := +((smooth_inv I).of_le le_top).cont_mdiff_at.cont_mdiff_within_at.comp x₀ hf $ set.maps_to_univ _ _ + +@[to_additive] +lemma cont_mdiff_at.inv {f : M → G} {x₀ : M} + (hf : cont_mdiff_at I' I n f x₀) : cont_mdiff_at I' I n (λx, (f x)⁻¹) x₀ := +((smooth_inv I).of_le le_top).cont_mdiff_at.comp x₀ hf + +@[to_additive] +lemma cont_mdiff_on.inv {f : M → G} {s : set M} + (hf : cont_mdiff_on I' I n f s) : cont_mdiff_on I' I n (λx, (f x)⁻¹) s := +λ x hx, (hf x hx).inv + +@[to_additive] +lemma cont_mdiff.inv {f : M → G} + (hf : cont_mdiff I' I n f) : cont_mdiff I' I n (λx, (f x)⁻¹) := +λ x, (hf x).inv + +@[to_additive] +lemma smooth_within_at.inv {f : M → G} {s : set M} {x₀ : M} + (hf : smooth_within_at I' I f s x₀) : smooth_within_at I' I (λx, (f x)⁻¹) s x₀ := +hf.inv + +@[to_additive] +lemma smooth_at.inv {f : M → G} {x₀ : M} + (hf : smooth_at I' I f x₀) : smooth_at I' I (λx, (f x)⁻¹) x₀ := +hf.inv @[to_additive] lemma smooth_on.inv {f : M → G} {s : set M} (hf : smooth_on I' I f s) : smooth_on I' I (λx, (f x)⁻¹) s := -(smooth_inv I).comp_smooth_on hf +hf.inv @[to_additive] -lemma smooth.div {f g : M → G} - (hf : smooth I' I f) (hg : smooth I' I g) : smooth I' I (f / g) := -by { rw div_eq_mul_inv, exact ((smooth_mul I).comp (hf.prod_mk hg.inv) : _), } +lemma smooth.inv {f : M → G} + (hf : smooth I' I f) : smooth I' I (λx, (f x)⁻¹) := +hf.inv + +@[to_additive] +lemma cont_mdiff_within_at.div {f g : M → G} {s : set M} {x₀ : M} + (hf : cont_mdiff_within_at I' I n f s x₀) (hg : cont_mdiff_within_at I' I n g s x₀) : + cont_mdiff_within_at I' I n (λ x, f x / g x) s x₀ := +by { simp_rw div_eq_mul_inv, exact hf.mul hg.inv } + +@[to_additive] +lemma cont_mdiff_at.div {f g : M → G} {x₀ : M} + (hf : cont_mdiff_at I' I n f x₀) (hg : cont_mdiff_at I' I n g x₀) : + cont_mdiff_at I' I n (λ x, f x / g x) x₀ := +by { simp_rw div_eq_mul_inv, exact hf.mul hg.inv } + +@[to_additive] +lemma cont_mdiff_on.div {f g : M → G} {s : set M} + (hf : cont_mdiff_on I' I n f s) (hg : cont_mdiff_on I' I n g s) : + cont_mdiff_on I' I n (λ x, f x / g x) s := +by { simp_rw div_eq_mul_inv, exact hf.mul hg.inv } + +@[to_additive] +lemma cont_mdiff.div {f g : M → G} + (hf : cont_mdiff I' I n f) (hg : cont_mdiff I' I n g) : + cont_mdiff I' I n (λ x, f x / g x) := +by { simp_rw div_eq_mul_inv, exact hf.mul hg.inv } + +@[to_additive] +lemma smooth_within_at.div {f g : M → G} {s : set M} {x₀ : M} + (hf : smooth_within_at I' I f s x₀) (hg : smooth_within_at I' I g s x₀) : + smooth_within_at I' I (λ x, f x / g x) s x₀ := +hf.div hg + +@[to_additive] +lemma smooth_at.div {f g : M → G} {x₀ : M} + (hf : smooth_at I' I f x₀) (hg : smooth_at I' I g x₀) : + smooth_at I' I (λ x, f x / g x) x₀ := +hf.div hg @[to_additive] lemma smooth_on.div {f g : M → G} {s : set M} (hf : smooth_on I' I f s) (hg : smooth_on I' I g s) : smooth_on I' I (f / g) s := -by { rw div_eq_mul_inv, exact ((smooth_mul I).comp_smooth_on (hf.prod_mk hg.inv) : _), } +hf.div hg + +@[to_additive] +lemma smooth.div {f g : M → G} + (hf : smooth I' I f) (hg : smooth I' I g) : smooth I' I (f / g) := +hf.div hg end lie_group diff --git a/src/geometry/manifold/vector_bundle/basic.lean b/src/geometry/manifold/vector_bundle/basic.lean index 01fc29ef7ebf3..5bef774d46b7f 100644 --- a/src/geometry/manifold/vector_bundle/basic.lean +++ b/src/geometry/manifold/vector_bundle/basic.lean @@ -174,6 +174,12 @@ lemma cont_mdiff_at_total_space (f : M → total_space E) (x₀ : M) : cont_mdiff_at IM 𝓘(𝕜, F) n (λ x, (trivialization_at F E (f x₀).proj (f x)).2) x₀ := by { simp_rw [← cont_mdiff_within_at_univ], exact cont_mdiff_within_at_total_space f } +/-- Characterization of C^n sections of a smooth vector bundle. -/ +lemma cont_mdiff_at_section (s : Π x, E x) (x₀ : B) : + cont_mdiff_at IB (IB.prod (𝓘(𝕜, F))) n (λ x, total_space_mk x (s x)) x₀ ↔ + cont_mdiff_at IB 𝓘(𝕜, F) n (λ x, (trivialization_at F E x₀ (total_space_mk x (s x))).2) x₀ := +by { simp_rw [cont_mdiff_at_total_space, and_iff_right_iff_imp], intro x, exact cont_mdiff_at_id } + variables (E) lemma cont_mdiff_proj : cont_mdiff (IB.prod 𝓘(𝕜, F)) IB n (π E) := begin diff --git a/src/geometry/manifold/vector_bundle/smooth_section.lean b/src/geometry/manifold/vector_bundle/smooth_section.lean new file mode 100644 index 0000000000000..9190ea480622e --- /dev/null +++ b/src/geometry/manifold/vector_bundle/smooth_section.lean @@ -0,0 +1,210 @@ +/- +Copyright © 2023 Heather Macbeth. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Heather Macbeth, Floris van Doorn +-/ + +import geometry.manifold.cont_mdiff_mfderiv +import topology.continuous_function.basic +import geometry.manifold.algebra.lie_group + +/-! +# Smooth sections + +In this file we define the type `cont_mdiff_section` of `n` times continuously differentiable +sections of a smooth vector bundle over a manifold `M` and prove that it's a module. +-/ +open bundle filter function +open_locale manifold + +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] +{E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] +{E' : Type*} [normed_add_comm_group E'] [normed_space 𝕜 E'] +{H : Type*} [topological_space H] +{H' : Type*} [topological_space H'] +(I : model_with_corners 𝕜 E H) (I' : model_with_corners 𝕜 E' H') +{M : Type*} [topological_space M] [charted_space H M] +{M' : Type*} [topological_space M'] [charted_space H' M'] +{E'' : Type*} [normed_add_comm_group E''] [normed_space 𝕜 E''] +{H'' : Type*} [topological_space H''] +{I'' : model_with_corners 𝕜 E'' H''} +{M'' : Type*} [topological_space M''] [charted_space H'' M''] +[smooth_manifold_with_corners I M] + + +variables (F : Type*) [normed_add_comm_group F] [normed_space 𝕜 F] -- `F` model fiber + (n : ℕ∞) + (V : M → Type*) [topological_space (total_space V)] -- `V` vector bundle + [Π x, add_comm_group (V x)] [Π x, module 𝕜 (V x)] +variables [Π x : M, topological_space (V x)] + [fiber_bundle F V] + [vector_bundle 𝕜 F V] + [smooth_vector_bundle F V I] + +/-- Bundled `n` times continuously differentiable sections of a vector bundle. -/ +@[protect_proj] +structure cont_mdiff_section := +(to_fun : Π x, V x) +(cont_mdiff_to_fun : cont_mdiff I (I.prod 𝓘(𝕜, F)) n (λ x, total_space_mk x (to_fun x))) + +/-- Bundled smooth sections of a vector bundle. -/ +@[reducible] def smooth_section := cont_mdiff_section I F ⊤ V + +localized "notation (name := cont_mdiff_section) `Cₛ^` n `⟮` I `; ` F `, ` V `⟯` := + cont_mdiff_section I F n V" in manifold + +namespace cont_mdiff_section + +variables {I} {I'} {n} {F} {V} + +instance : has_coe_to_fun Cₛ^n⟮I; F, V⟯ (λ s, Π x, V x) := ⟨cont_mdiff_section.to_fun⟩ + +variables {s t : Cₛ^n⟮I; F, V⟯} + +@[simp] lemma coe_fn_mk (s : Π x, V x) + (hs : cont_mdiff I (I.prod 𝓘(𝕜, F)) n (λ x, total_space_mk x (s x))) : + (mk s hs : Π x, V x) = s := +rfl + +protected lemma cont_mdiff (s : Cₛ^n⟮I; F, V⟯) : + cont_mdiff I (I.prod 𝓘(𝕜, F)) n (λ x, total_space_mk x (s x : V x)) := s.cont_mdiff_to_fun + +protected lemma smooth (s : Cₛ^∞⟮I; F, V⟯) : + smooth I (I.prod 𝓘(𝕜, F)) (λ x, total_space_mk x (s x : V x)) := s.cont_mdiff_to_fun + +protected lemma mdifferentiable' (s : Cₛ^n⟮I; F, V⟯) (hn : 1 ≤ n) : + mdifferentiable I (I.prod 𝓘(𝕜, F)) (λ x, total_space_mk x (s x : V x)) := +s.cont_mdiff.mdifferentiable hn + +protected lemma mdifferentiable (s : Cₛ^∞⟮I; F, V⟯) : + mdifferentiable I (I.prod 𝓘(𝕜, F)) (λ x, total_space_mk x (s x : V x)) := +s.cont_mdiff.mdifferentiable le_top + +protected lemma mdifferentiable_at (s : Cₛ^∞⟮I; F, V⟯) {x} : + mdifferentiable_at I (I.prod 𝓘(𝕜, F)) (λ x, total_space_mk x (s x : V x)) x := +s.mdifferentiable x + +lemma coe_inj ⦃s t : Cₛ^n⟮I; F, V⟯⦄ (h : (s : Π x, V x) = t) : s = t := +by cases s; cases t; cases h; refl + +lemma coe_injective : injective (coe_fn : Cₛ^n⟮I; F, V⟯ → Π x, V x) := +coe_inj + +@[ext] theorem ext (h : ∀ x, s x = t x) : s = t := +by cases s; cases t; congr'; exact funext h + +instance has_add : has_add Cₛ^n⟮I; F, V⟯ := +begin + refine ⟨λ s t, ⟨s + t, _⟩⟩, + intro x₀, + have hs := s.cont_mdiff x₀, + have ht := t.cont_mdiff x₀, + rw [cont_mdiff_at_section] at hs ht ⊢, + set e := trivialization_at F V x₀, + refine (hs.add ht).congr_of_eventually_eq _, + refine eventually_of_mem (e.open_base_set.mem_nhds $ mem_base_set_trivialization_at F V x₀) _, + intros x hx, + apply (e.linear 𝕜 hx).1, +end + +@[simp] +lemma coe_add (s t : Cₛ^n⟮I; F, V⟯) : ⇑(s + t) = s + t := rfl + +instance has_sub : has_sub Cₛ^n⟮I; F, V⟯ := +begin + refine ⟨λ s t, ⟨s - t, _⟩⟩, + intro x₀, + have hs := s.cont_mdiff x₀, + have ht := t.cont_mdiff x₀, + rw [cont_mdiff_at_section] at hs ht ⊢, + set e := trivialization_at F V x₀, + refine (hs.sub ht).congr_of_eventually_eq _, + refine eventually_of_mem (e.open_base_set.mem_nhds $ mem_base_set_trivialization_at F V x₀) _, + intros x hx, + apply (e.linear 𝕜 hx).map_sub, +end + +@[simp] +lemma coe_sub (s t : Cₛ^n⟮I; F, V⟯) : ⇑(s - t) = s - t := rfl + +instance has_zero : has_zero Cₛ^n⟮I; F, V⟯ := +⟨⟨λ x, 0, (smooth_zero_section 𝕜 V).of_le le_top⟩⟩ + +instance inhabited : inhabited Cₛ^n⟮I; F, V⟯ := ⟨0⟩ + +@[simp] +lemma coe_zero : ⇑(0 : Cₛ^n⟮I; F, V⟯) = 0 := rfl + +instance has_smul : has_smul 𝕜 Cₛ^n⟮I; F, V⟯ := +begin + refine ⟨λ c s, ⟨c • s, _⟩⟩, + intro x₀, + have hs := s.cont_mdiff x₀, + rw [cont_mdiff_at_section] at hs ⊢, + set e := trivialization_at F V x₀, + refine (cont_mdiff_at_const.smul hs).congr_of_eventually_eq _, + { exact c }, + refine eventually_of_mem (e.open_base_set.mem_nhds $ mem_base_set_trivialization_at F V x₀) _, + intros x hx, + apply (e.linear 𝕜 hx).2, +end + +@[simp] +lemma coe_smul (r : 𝕜) (s : Cₛ^n⟮I; F, V⟯) : ⇑(r • s : Cₛ^n⟮I; F, V⟯) = r • s := rfl + +instance has_neg : has_neg Cₛ^n⟮I; F, V⟯ := +begin + refine ⟨λ s, ⟨- s, _⟩⟩, + intro x₀, + have hs := s.cont_mdiff x₀, + rw [cont_mdiff_at_section] at hs ⊢, + set e := trivialization_at F V x₀, + refine hs.neg.congr_of_eventually_eq _, + refine eventually_of_mem (e.open_base_set.mem_nhds $ mem_base_set_trivialization_at F V x₀) _, + intros x hx, + apply (e.linear 𝕜 hx).map_neg +end + +@[simp] +lemma coe_neg (s : Cₛ^n⟮I; F, V⟯) : ⇑(- s : Cₛ^n⟮I; F, V⟯) = - s := rfl + +instance has_nsmul : has_smul ℕ Cₛ^n⟮I; F, V⟯ := +⟨nsmul_rec⟩ + +@[simp] +lemma coe_nsmul (s : Cₛ^n⟮I; F, V⟯) (k : ℕ) : ⇑(k • s : Cₛ^n⟮I; F, V⟯) = k • s := +begin + induction k with k ih, + { simp_rw [zero_smul], refl }, + simp_rw [succ_nsmul, ← ih], refl, +end + +instance has_zsmul : has_smul ℤ Cₛ^n⟮I; F, V⟯ := +⟨zsmul_rec⟩ + +@[simp] +lemma coe_zsmul (s : Cₛ^n⟮I; F, V⟯) (z : ℤ) : ⇑(z • s : Cₛ^n⟮I; F, V⟯) = z • s := +begin + cases z with n n, + refine (coe_nsmul s n).trans _, + simp only [int.of_nat_eq_coe, coe_nat_zsmul], + refine (congr_arg has_neg.neg (coe_nsmul s (n+1))).trans _, + simp only [zsmul_neg_succ_of_nat, neg_inj] +end + +instance add_comm_group : add_comm_group Cₛ^n⟮I; F, V⟯ := +coe_injective.add_comm_group _ coe_zero coe_add coe_neg coe_sub coe_nsmul coe_zsmul + +variables (I F V n) +/-- The additive morphism from smooth sections to dependent maps. -/ +def coe_add_hom : Cₛ^n⟮I; F, V⟯ →+ Π x, V x := +{ to_fun := coe_fn, + map_zero' := coe_zero, + map_add' := coe_add } + +variables {I F V n} + +instance module : module 𝕜 Cₛ^n⟮I; F, V⟯ := +coe_injective.module 𝕜 (coe_add_hom I F n V) coe_smul + +end cont_mdiff_section From 8bdf5e9b7f395d20e104dd1e309316daee8ad547 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 14 Jun 2023 13:43:54 +0000 Subject: [PATCH 1185/1291] refactor(field_theory/polynomial_galois_group): remove `open_locale classical` (#19184) This doesn't actually generalize any lemmas, but it reduces the chance of lemmas added in future being less general. The `restrict_dvd_def` lemma is new, and deals with substituting an arbitray `decidable_eq` instance. --- src/field_theory/polynomial_galois_group.lean | 17 +++++++++++++---- 1 file changed, 13 insertions(+), 4 deletions(-) diff --git a/src/field_theory/polynomial_galois_group.lean b/src/field_theory/polynomial_galois_group.lean index ad3d3e5a0757e..76a94be444e84 100644 --- a/src/field_theory/polynomial_galois_group.lean +++ b/src/field_theory/polynomial_galois_group.lean @@ -39,7 +39,7 @@ equals the number of real roots plus the number of roots not fixed by complex co -/ noncomputable theory -open_locale classical polynomial +open_locale polynomial open finite_dimensional @@ -139,6 +139,7 @@ begin have hy := subtype.mem y, simp only [root_set, finset.mem_coe, multiset.mem_to_finset, key, multiset.mem_map] at hy, rcases hy with ⟨x, hx1, hx2⟩, + classical, exact ⟨⟨x, multiset.mem_to_finset.mpr hx1⟩, subtype.ext hx2⟩ } end @@ -201,12 +202,18 @@ variables {p q} /-- `polynomial.gal.restrict`, when both fields are splitting fields of polynomials. -/ def restrict_dvd (hpq : p ∣ q) : q.gal →* p.gal := +by haveI := classical.dec (q = 0); exact if hq : q = 0 then 1 else @restrict F _ p _ _ _ ⟨splits_of_splits_of_dvd (algebra_map F q.splitting_field) hq (splitting_field.splits q) hpq⟩ +lemma restrict_dvd_def [decidable (q = 0)] (hpq : p ∣ q) : + restrict_dvd hpq = if hq : q = 0 then 1 else @restrict F _ p _ _ _ + ⟨splits_of_splits_of_dvd (algebra_map F q.splitting_field) hq (splitting_field.splits q) hpq⟩ := +by convert rfl + lemma restrict_dvd_surjective (hpq : p ∣ q) (hq : q ≠ 0) : function.surjective (restrict_dvd hpq) := -by simp only [restrict_dvd, dif_neg hq, restrict_surjective] +by classical; simp only [restrict_dvd_def, dif_neg hq, restrict_surjective] variables (p q) @@ -221,10 +228,11 @@ begin { haveI : unique (p * q).gal, { rw hpq, apply_instance }, exact λ f g h, eq.trans (unique.eq_default f) (unique.eq_default g).symm }, intros f g hfg, - dsimp only [restrict_prod, restrict_dvd] at hfg, + classical, + simp only [restrict_prod, restrict_dvd_def] at hfg, simp only [dif_neg hpq, monoid_hom.prod_apply, prod.mk.inj_iff] at hfg, ext x hx, - rw [root_set, polynomial.map_mul, polynomial.roots_mul] at hx, + rw [root_set_def, polynomial.map_mul, polynomial.roots_mul] at hx, cases multiset.mem_add.mp (multiset.mem_to_finset.mp hx) with h h, { haveI : fact (p.splits (algebra_map F (p * q).splitting_field)) := ⟨splits_of_splits_of_dvd _ hpq (splitting_field.splits (p * q)) (dvd_mul_right p q)⟩, @@ -401,6 +409,7 @@ lemma gal_action_hom_bijective_of_prime_degree (p_roots : fintype.card (p.root_set ℂ) = fintype.card (p.root_set ℝ) + 2) : function.bijective (gal_action_hom p ℂ) := begin + classical, have h1 : fintype.card (p.root_set ℂ) = p.nat_degree, { simp_rw [root_set_def, finset.coe_sort_coe, fintype.card_coe], rw [multiset.to_finset_card_of_nodup, ←nat_degree_eq_card_roots], From c8734e8953e4b439147bd6f75c2163f6d27cdce6 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 14 Jun 2023 13:43:55 +0000 Subject: [PATCH 1186/1291] feat(data/mv_polynomial/basic): add more general smul compatibility instances (#19187) These apply for things like actions by the units of the base ring. --- src/data/mv_polynomial/basic.lean | 9 ++++++--- 1 file changed, 6 insertions(+), 3 deletions(-) diff --git a/src/data/mv_polynomial/basic.lean b/src/data/mv_polynomial/basic.lean index 40641f917b012..8dac8e5c459bb 100644 --- a/src/data/mv_polynomial/basic.lean +++ b/src/data/mv_polynomial/basic.lean @@ -130,10 +130,13 @@ add_monoid_algebra.is_central_scalar instance [comm_semiring R] [comm_semiring S₁] [algebra R S₁] : algebra R (mv_polynomial σ S₁) := add_monoid_algebra.algebra --- Register with high priority to avoid timeout in `data.mv_polynomial.pderiv` -instance is_scalar_tower' [comm_semiring R] [comm_semiring S₁] [algebra R S₁] : +instance is_scalar_tower_right [comm_semiring S₁] [distrib_smul R S₁] [is_scalar_tower R S₁ S₁] : is_scalar_tower R (mv_polynomial σ S₁) (mv_polynomial σ S₁) := -is_scalar_tower.right +add_monoid_algebra.is_scalar_tower_self _ + +instance smul_comm_class_right [comm_semiring S₁] [distrib_smul R S₁] [smul_comm_class R S₁ S₁] : + smul_comm_class R (mv_polynomial σ S₁) (mv_polynomial σ S₁) := +add_monoid_algebra.smul_comm_class_self _ /-- If `R` is a subsingleton, then `mv_polynomial σ R` has a unique element -/ instance unique [comm_semiring R] [subsingleton R] : unique (mv_polynomial σ R) := From 9f55d0d4363ae59948c33864cbc52e0b12e0e8ce Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Wed, 14 Jun 2023 16:17:48 +0000 Subject: [PATCH 1187/1291] feat(measure_theory/constructions/polish): quotient group is a Borel space (#19186) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit * Suslin's theorem: an analytic set with analytic complement is measurable. * Image of a measurable set in a Polish space under a measurable map is an analytic set. * Preimage of a set under a measurable surjective map from a Polish space is measurable iff the original set is measurable. * Quotient space of a Polish space with quotient σ-algebra is a Borel space provided that it has second countable topology. * In particular, quotient group of a Polish topological group is a Borel space. * Change instance for `measurable_space` on `add_circle`. --- .../constructions/borel_space/basic.lean | 3 + src/measure_theory/constructions/polish.lean | 175 ++++++++++++++++-- src/measure_theory/integral/periodic.lean | 8 +- 3 files changed, 169 insertions(+), 17 deletions(-) diff --git a/src/measure_theory/constructions/borel_space/basic.lean b/src/measure_theory/constructions/borel_space/basic.lean index 73cbb8b06bdc0..95c59b3ffd372 100644 --- a/src/measure_theory/constructions/borel_space/basic.lean +++ b/src/measure_theory/constructions/borel_space/basic.lean @@ -58,6 +58,9 @@ open measurable_space topological_space def borel (α : Type u) [topological_space α] : measurable_space α := generate_from {s : set α | is_open s} +theorem borel_anti : antitone (@borel α) := +λ _ _ h, measurable_space.generate_from_le $ λ s hs, generate_measurable.basic _ (h _ hs) + lemma borel_eq_top_of_discrete [topological_space α] [discrete_topology α] : borel α = ⊤ := top_le_iff.1 $ λ s hs, generate_measurable.basic s (is_open_discrete s) diff --git a/src/measure_theory/constructions/polish.lean b/src/measure_theory/constructions/polish.lean index 720a7aac73281..63cf523501aca 100644 --- a/src/measure_theory/constructions/polish.lean +++ b/src/measure_theory/constructions/polish.lean @@ -253,24 +253,42 @@ end a finer Polish topology on the source space for which the function is continuous. -/ lemma _root_.measurable.exists_continuous {α β : Type*} [t : topological_space α] [polish_space α] [measurable_space α] [borel_space α] - [tβ : topological_space β] [second_countable_topology β] [measurable_space β] [borel_space β] - {f : α → β} (hf : measurable f) : + [tβ : topological_space β] [measurable_space β] [opens_measurable_space β] + {f : α → β} [second_countable_topology (range f)] (hf : measurable f) : ∃ (t' : topological_space α), t' ≤ t ∧ @continuous α β t' tβ f ∧ @polish_space α t' := begin - obtain ⟨b, b_count, -, hb⟩ : ∃b : set (set β), b.countable ∧ ∅ ∉ b ∧ is_topological_basis b := - exists_countable_basis β, - haveI : encodable b := b_count.to_encodable, - have : ∀ (s : b), is_clopenable (f ⁻¹' s), + obtain ⟨b, b_count, -, hb⟩ : + ∃ b : set (set (range f)), b.countable ∧ ∅ ∉ b ∧ is_topological_basis b := + exists_countable_basis (range f), + haveI : countable b := b_count.to_subtype, + have : ∀ (s : b), is_clopenable (range_factorization f ⁻¹' s), { assume s, apply measurable_set.is_clopenable, - exact hf (hb.is_open s.2).measurable_set }, + exact hf.subtype_mk (hb.is_open s.2).measurable_set }, choose T Tt Tpolish Tclosed Topen using this, obtain ⟨t', t'T, t't, t'_polish⟩ : ∃ (t' : topological_space α), (∀ i, t' ≤ T i) ∧ (t' ≤ t) ∧ @polish_space α t' := exists_polish_space_forall_le T Tt Tpolish, + letI := t', -- not needed in Lean 4 refine ⟨t', t't, _, t'_polish⟩, - apply hb.continuous _ (λ s hs, _), - exact t'T ⟨s, hs⟩ _ (Topen ⟨s, hs⟩), + have : @continuous _ _ t' _ (range_factorization f) := + hb.continuous _ (λ s hs, t'T ⟨s, hs⟩ _ (Topen ⟨s, hs⟩)), + exact continuous_subtype_coe.comp this +end + +/-- The image of a measurable set in a Polish space under a measurable map is an analytic set. -/ +theorem _root_.measurable_set.analytic_set_image {X Y : Type*} + [topological_space X] [polish_space X] [measurable_space X] [borel_space X] + [topological_space Y] [measurable_space Y] [opens_measurable_space Y] + {f : X → Y} [second_countable_topology (range f)] {s : set X} (hs : measurable_set s) + (hf : measurable f) : analytic_set (f '' s) := +begin + borelize X, + rcases hf.exists_continuous with ⟨τ', hle, hfc, hτ'⟩, + letI m' : measurable_space X := @borel _ τ', + haveI b' : borel_space X := ⟨rfl⟩, + have hle := borel_anti hle, + exact (hle _ hs).analytic_set.image_of_continuous hfc end /-! ### Separating sets with measurable sets -/ @@ -302,8 +320,9 @@ end contained in disjoint Borel sets (see the full statement in `analytic_set.measurably_separable`). Here, we prove this when our analytic sets are the ranges of functions from `ℕ → ℕ`. -/ -lemma measurably_separable_range_of_disjoint [t2_space α] [measurable_space α] [borel_space α] - {f g : (ℕ → ℕ) → α} (hf : continuous f) (hg : continuous g) (h : disjoint (range f) (range g)) : +lemma measurably_separable_range_of_disjoint [t2_space α] [measurable_space α] + [opens_measurable_space α] {f g : (ℕ → ℕ) → α} (hf : continuous f) (hg : continuous g) + (h : disjoint (range f) (range g)) : measurably_separable (range f) (range g) := begin /- We follow [Kechris, *Classical Descriptive Set Theory* (Theorem 14.7)][kechris1995]. @@ -416,8 +435,9 @@ end /-- The Lusin separation theorem: if two analytic sets are disjoint, then they are contained in disjoint Borel sets. -/ -theorem analytic_set.measurably_separable [t2_space α] [measurable_space α] [borel_space α] - {s t : set α} (hs : analytic_set s) (ht : analytic_set t) (h : disjoint s t) : +theorem analytic_set.measurably_separable [t2_space α] [measurable_space α] + [opens_measurable_space α] {s t : set α} (hs : analytic_set s) (ht : analytic_set t) + (h : disjoint s t) : measurably_separable s t := begin rw analytic_set at hs ht, @@ -428,6 +448,135 @@ begin exact measurably_separable_range_of_disjoint f_cont g_cont h, end +/-- **Suslin's Theorem**: in a Hausdorff topological space, an analytic set with an analytic +complement is measurable. -/ +theorem analytic_set.measurable_set_of_compl [t2_space α] [measurable_space α] + [opens_measurable_space α] {s : set α} (hs : analytic_set s) (hsc : analytic_set (sᶜ)) : + measurable_set s := +begin + rcases hs.measurably_separable hsc disjoint_compl_right with ⟨u, hsu, hdu, hmu⟩, + obtain rfl : s = u := hsu.antisymm (disjoint_compl_left_iff_subset.1 hdu), + exact hmu +end + +end measure_theory + +/-! +### Measurability of preimages under measurable maps +-/ + +namespace measurable + +variables {X Y β : Type*} + [topological_space X] [polish_space X] [measurable_space X] [borel_space X] + [topological_space Y] [t2_space Y] [measurable_space Y] [opens_measurable_space Y] + [measurable_space β] + +/-- If `f : X → Y` is a surjective Borel measurable map from a Polish space to a topological space +with second countable topology, then the preimage of a set `s` is measurable if and only if the set +is measurable. +One implication is the definition of measurability, the other one heavily relies on `X` being a +Polish space. -/ +theorem measurable_set_preimage_iff_of_surjective [second_countable_topology Y] {f : X → Y} + (hf : measurable f) (hsurj : surjective f) {s : set Y} : + measurable_set (f ⁻¹' s) ↔ measurable_set s := +begin + refine ⟨λ h, _, λ h, hf h⟩, + apply measure_theory.analytic_set.measurable_set_of_compl, + { rw [← image_preimage_eq s hsurj], + exact h.analytic_set_image hf }, + { rw [← image_preimage_eq (sᶜ) hsurj], + exact h.compl.analytic_set_image hf } +end + +theorem map_measurable_space_eq [second_countable_topology Y] {f : X → Y} (hf : measurable f) + (hsurj : surjective f) : measurable_space.map f ‹measurable_space X› = ‹measurable_space Y› := +measurable_space.ext $ λ _, hf.measurable_set_preimage_iff_of_surjective hsurj + +theorem map_measurable_space_eq_borel [second_countable_topology Y] {f : X → Y} (hf : measurable f) + (hsurj : surjective f) : measurable_space.map f ‹measurable_space X› = borel Y := +begin + have := hf.mono le_rfl opens_measurable_space.borel_le, + letI := borel Y, haveI : borel_space Y := ⟨rfl⟩, + exact this.map_measurable_space_eq hsurj +end + +theorem borel_space_codomain [second_countable_topology Y] {f : X → Y} (hf : measurable f) + (hsurj : surjective f) : borel_space Y := +⟨(hf.map_measurable_space_eq hsurj).symm.trans $ hf.map_measurable_space_eq_borel hsurj⟩ + +/-- If `f : X → Y` is a Borel measurable map from a Polish space to a topological space with second +countable topology, then the preimage of a set `s` is measurable if and only if the set is +measurable in `set.range f`. -/ +theorem measurable_set_preimage_iff_preimage_coe {f : X → Y} [second_countable_topology (range f)] + (hf : measurable f) {s : set Y} : + measurable_set (f ⁻¹' s) ↔ measurable_set (coe ⁻¹' s : set (range f)) := +have hf' : measurable (range_factorization f) := hf.subtype_mk, +by rw [← hf'.measurable_set_preimage_iff_of_surjective surjective_onto_range]; refl + +/-- If `f : X → Y` is a Borel measurable map from a Polish space to a topological space with second +countable topology and the range of `f` is measurable, then the preimage of a set `s` is measurable +if and only if the intesection with `set.range f` is measurable. -/ +theorem measurable_set_preimage_iff_inter_range {f : X → Y} [second_countable_topology (range f)] + (hf : measurable f) (hr : measurable_set (range f)) {s : set Y} : + measurable_set (f ⁻¹' s) ↔ measurable_set (s ∩ range f) := +begin + rw [hf.measurable_set_preimage_iff_preimage_coe, + ← (measurable_embedding.subtype_coe hr).measurable_set_image, subtype.image_preimage_coe] +end + +/-- If `f : X → Y` is a Borel measurable map from a Polish space to a topological space with second +countable topology, then for any measurable space `β` and `g : Y → β`, the composition `g ∘ f` is +measurable if and only if the restriction of `g` to the range of `f` is measurable. -/ +theorem measurable_comp_iff_restrict {f : X → Y} [second_countable_topology (range f)] + (hf : measurable f) {g : Y → β} : + measurable (g ∘ f) ↔ measurable (restrict (range f) g) := +forall₂_congr $ λ s _, + @measurable.measurable_set_preimage_iff_preimage_coe _ _ _ _ _ _ _ _ _ _ _ _ hf (g ⁻¹' s) + +/-- If `f : X → Y` is a surjective Borel measurable map from a Polish space to a topological space +with second countable topology, then for any measurable space `α` and `g : Y → α`, the composition +`g ∘ f` is measurable if and only if `g` is measurable. -/ +theorem measurable_comp_iff_of_surjective [second_countable_topology Y] {f : X → Y} + (hf : measurable f) (hsurj : surjective f) {g : Y → β} : + measurable (g ∘ f) ↔ measurable g := +forall₂_congr $ λ s _, + @measurable.measurable_set_preimage_iff_of_surjective _ _ _ _ _ _ _ _ _ _ _ _ hf hsurj (g ⁻¹' s) + +end measurable + +theorem continuous.map_eq_borel {X Y : Type*} + [topological_space X] [polish_space X] [measurable_space X] [borel_space X] + [topological_space Y] [t2_space Y] [second_countable_topology Y] + {f : X → Y} (hf : continuous f) (hsurj : surjective f) : + measurable_space.map f ‹measurable_space X› = borel Y := +begin + borelize Y, + exact hf.measurable.map_measurable_space_eq hsurj +end + +theorem continuous.map_borel_eq {X Y : Type*} [topological_space X] [polish_space X] + [topological_space Y] [t2_space Y] [second_countable_topology Y] + {f : X → Y} (hf : continuous f) (hsurj : surjective f) : + measurable_space.map f (borel X) = borel Y := +begin + borelize X, + exact hf.map_eq_borel hsurj +end + +instance quotient.borel_space {X : Type*} [topological_space X] [polish_space X] + [measurable_space X] [borel_space X] {s : setoid X} [t2_space (quotient s)] + [second_countable_topology (quotient s)] : borel_space (quotient s) := +⟨continuous_quotient_mk.map_eq_borel (surjective_quotient_mk _)⟩ + +@[to_additive] +instance quotient_group.borel_space {G : Type*} [topological_space G] [polish_space G] + [group G] [topological_group G] [measurable_space G] [borel_space G] + {N : subgroup G} [N.normal] [is_closed (N : set G)] : borel_space (G ⧸ N) := +by haveI := polish_space.second_countable G; exact quotient.borel_space + +namespace measure_theory + /-! ### Injective images of Borel sets -/ variables {γ : Type*} [tγ : topological_space γ] [polish_space γ] diff --git a/src/measure_theory/integral/periodic.lean b/src/measure_theory/integral/periodic.lean index 9c87a4e07a0bc..f34bc3b914757 100644 --- a/src/measure_theory/integral/periodic.lean +++ b/src/measure_theory/integral/periodic.lean @@ -5,6 +5,7 @@ Authors: Yury Kudryashov, Alex Kontorovich, Heather Macbeth -/ import measure_theory.measure.lebesgue.eq_haar import measure_theory.measure.haar.quotient +import measure_theory.constructions.polish import measure_theory.integral.interval_integral import topology.algebra.order.floor @@ -28,12 +29,11 @@ open set function measure_theory measure_theory.measure topological_space add_su open_locale measure_theory nnreal ennreal -local attribute [-instance] quotient_add_group.measurable_space quotient.measurable_space - noncomputable instance add_circle.measurable_space {a : ℝ} : measurable_space (add_circle a) := -borel (add_circle a) +quotient_add_group.measurable_space _ -instance add_circle.borel_space {a : ℝ} : borel_space (add_circle a) := ⟨rfl⟩ +instance add_circle.borel_space {a : ℝ} : borel_space (add_circle a) := +quotient_add_group.borel_space @[measurability] protected lemma add_circle.measurable_mk' {a : ℝ} : measurable (coe : ℝ → add_circle a) := From 8af7091a43227e179939ba132e54e54e9f3b089a Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Thu, 15 Jun 2023 05:00:11 +0000 Subject: [PATCH 1188/1291] chore(*): add mathlib4 synchronization comments (#19190) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.homology.Module` * `algebra.lie.weights` * `algebraic_geometry.Scheme` * `algebraic_geometry.Spec` * `analysis.analytic.isolated_zeros` * `analysis.analytic.uniqueness` * `analysis.complex.open_mapping` * `analysis.mellin_transform` * `analysis.special_functions.gamma.basic` * `category_theory.monoidal.CommMon_` * `geometry.euclidean.angle.oriented.affine` * `geometry.euclidean.angle.oriented.right_angle` * `geometry.euclidean.angle.oriented.rotation` * `geometry.euclidean.circumcenter` * `geometry.euclidean.sphere.ptolemy` * `geometry.euclidean.triangle` * `measure_theory.function.conditional_expectation.unique` * `order.category.BddDistLat` * `order.category.BddLat` * `order.category.CompleteLat` * `order.category.HeytAlg` * `probability.kernel.measurable_integral` * `probability.kernel.with_density` * `ring_theory.ring_hom.finite` * `topology.continuous_function.ideals` * `topology.gluing` --- src/algebra/homology/Module.lean | 3 +++ src/algebra/lie/weights.lean | 3 +++ src/algebraic_geometry/Scheme.lean | 3 +++ src/algebraic_geometry/Spec.lean | 3 +++ src/analysis/analytic/isolated_zeros.lean | 3 +++ src/analysis/analytic/uniqueness.lean | 3 +++ src/analysis/complex/open_mapping.lean | 3 +++ src/analysis/mellin_transform.lean | 3 +++ src/analysis/special_functions/gamma/basic.lean | 3 +++ src/category_theory/monoidal/CommMon_.lean | 3 +++ src/geometry/euclidean/angle/oriented/affine.lean | 3 +++ src/geometry/euclidean/angle/oriented/right_angle.lean | 3 +++ src/geometry/euclidean/angle/oriented/rotation.lean | 3 +++ src/geometry/euclidean/circumcenter.lean | 3 +++ src/geometry/euclidean/sphere/ptolemy.lean | 3 +++ src/geometry/euclidean/triangle.lean | 3 +++ .../function/conditional_expectation/unique.lean | 3 +++ src/order/category/BddDistLat.lean | 3 +++ src/order/category/BddLat.lean | 3 +++ src/order/category/CompleteLat.lean | 3 +++ src/order/category/HeytAlg.lean | 3 +++ src/probability/kernel/measurable_integral.lean | 3 +++ src/probability/kernel/with_density.lean | 3 +++ src/ring_theory/ring_hom/finite.lean | 3 +++ src/topology/continuous_function/ideals.lean | 3 +++ src/topology/gluing.lean | 3 +++ 26 files changed, 78 insertions(+) diff --git a/src/algebra/homology/Module.lean b/src/algebra/homology/Module.lean index 6e13beda5b292..b2a096d7dfe9e 100644 --- a/src/algebra/homology/Module.lean +++ b/src/algebra/homology/Module.lean @@ -11,6 +11,9 @@ import category_theory.limits.concrete_category /-! # Complexes of modules +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We provide some additional API to work with homological complexes in `Module R`. -/ diff --git a/src/algebra/lie/weights.lean b/src/algebra/lie/weights.lean index cfaced0c55890..32bc6b4540044 100644 --- a/src/algebra/lie/weights.lean +++ b/src/algebra/lie/weights.lean @@ -14,6 +14,9 @@ import ring_theory.tensor_product /-! # Weights and roots of Lie modules and Lie algebras +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Just as a key tool when studying the behaviour of a linear operator is to decompose the space on which it acts into a sum of (generalised) eigenspaces, a key tool when studying a representation `M` of Lie algebra `L` is to decompose `M` into a sum of simultaneous eigenspaces of `x` as `x` ranges diff --git a/src/algebraic_geometry/Scheme.lean b/src/algebraic_geometry/Scheme.lean index 7a76e19a7d62e..888440d98ab1e 100644 --- a/src/algebraic_geometry/Scheme.lean +++ b/src/algebraic_geometry/Scheme.lean @@ -9,6 +9,9 @@ import algebra.category.Ring.constructions /-! # The category of schemes +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A scheme is a locally ringed space such that every point is contained in some open set where there is an isomorphism of presheaves between the restriction to that open set, and the structure sheaf of `Spec R`, for some commutative ring `R`. diff --git a/src/algebraic_geometry/Spec.lean b/src/algebraic_geometry/Spec.lean index 8663af3575c76..eff89cbd23173 100644 --- a/src/algebraic_geometry/Spec.lean +++ b/src/algebraic_geometry/Spec.lean @@ -13,6 +13,9 @@ import algebra.module.localized_module /-! # $Spec$ as a functor to locally ringed spaces. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define the functor $Spec$ from commutative rings to locally ringed spaces. ## Implementation notes diff --git a/src/analysis/analytic/isolated_zeros.lean b/src/analysis/analytic/isolated_zeros.lean index 98343d71e3c9a..7759f8579ca8b 100644 --- a/src/analysis/analytic/isolated_zeros.lean +++ b/src/analysis/analytic/isolated_zeros.lean @@ -12,6 +12,9 @@ import analysis.analytic.uniqueness /-! # Principle of isolated zeros +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves the fact that the zeros of a non-constant analytic function of one variable are isolated. It also introduces a little bit of API in the `has_fpower_series_at` namespace that is useful in this setup. diff --git a/src/analysis/analytic/uniqueness.lean b/src/analysis/analytic/uniqueness.lean index 93dd668cc9339..8668e4da3b970 100644 --- a/src/analysis/analytic/uniqueness.lean +++ b/src/analysis/analytic/uniqueness.lean @@ -10,6 +10,9 @@ import analysis.normed_space.completion /-! # Uniqueness principle for analytic functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We show that two analytic functions which coincide around a point coincide on whole connected sets, in `analytic_on.eq_on_of_preconnected_of_eventually_eq`. -/ diff --git a/src/analysis/complex/open_mapping.lean b/src/analysis/complex/open_mapping.lean index 9e347d273ff1a..1aa8c27c29283 100644 --- a/src/analysis/complex/open_mapping.lean +++ b/src/analysis/complex/open_mapping.lean @@ -10,6 +10,9 @@ import analysis.complex.abs_max /-! # The open mapping theorem for holomorphic functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves the open mapping theorem for holomorphic functions, namely that an analytic function on a preconnected set of the complex plane is either constant or open. The main step is to show a local version of the theorem that states that if `f` is analytic at a point `z₀`, then either diff --git a/src/analysis/mellin_transform.lean b/src/analysis/mellin_transform.lean index c3aae84253868..ac6518d306041 100644 --- a/src/analysis/mellin_transform.lean +++ b/src/analysis/mellin_transform.lean @@ -10,6 +10,9 @@ import measure_theory.measure.haar.normed_space /-! # The Mellin transform +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define the Mellin transform of a locally integrable function on `Ioi 0`, and show it is differentiable in a suitable vertical strip. diff --git a/src/analysis/special_functions/gamma/basic.lean b/src/analysis/special_functions/gamma/basic.lean index 07381ea31b077..e8cdf1d2b63ba 100644 --- a/src/analysis/special_functions/gamma/basic.lean +++ b/src/analysis/special_functions/gamma/basic.lean @@ -10,6 +10,9 @@ import analysis.mellin_transform /-! # The Gamma function +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the `Γ` function (of a real or complex variable `s`). We define this by Euler's integral `Γ(s) = ∫ x in Ioi 0, exp (-x) * x ^ (s - 1)` in the range where this integral converges (i.e., for `0 < s` in the real case, and `0 < re s` in the complex case). diff --git a/src/category_theory/monoidal/CommMon_.lean b/src/category_theory/monoidal/CommMon_.lean index 949b2ac4e0357..b772938e4888b 100644 --- a/src/category_theory/monoidal/CommMon_.lean +++ b/src/category_theory/monoidal/CommMon_.lean @@ -8,6 +8,9 @@ import category_theory.monoidal.Mon_ /-! # The category of commutative monoids in a braided monoidal category. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ universes v₁ v₂ u₁ u₂ u diff --git a/src/geometry/euclidean/angle/oriented/affine.lean b/src/geometry/euclidean/angle/oriented/affine.lean index 96caff2698604..4539469cf2dad 100644 --- a/src/geometry/euclidean/angle/oriented/affine.lean +++ b/src/geometry/euclidean/angle/oriented/affine.lean @@ -10,6 +10,9 @@ import geometry.euclidean.angle.unoriented.affine /-! # Oriented angles. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines oriented angles in Euclidean affine spaces. ## Main definitions diff --git a/src/geometry/euclidean/angle/oriented/right_angle.lean b/src/geometry/euclidean/angle/oriented/right_angle.lean index a245fc84aa0d4..b38ec3b422bf2 100644 --- a/src/geometry/euclidean/angle/oriented/right_angle.lean +++ b/src/geometry/euclidean/angle/oriented/right_angle.lean @@ -9,6 +9,9 @@ import geometry.euclidean.angle.unoriented.right_angle /-! # Oriented angles in right-angled triangles. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves basic geometrical results about distances and oriented angles in (possibly degenerate) right-angled triangles in real inner product spaces and Euclidean affine spaces. diff --git a/src/geometry/euclidean/angle/oriented/rotation.lean b/src/geometry/euclidean/angle/oriented/rotation.lean index 890777df7fb5f..8ca7a0bcff940 100644 --- a/src/geometry/euclidean/angle/oriented/rotation.lean +++ b/src/geometry/euclidean/angle/oriented/rotation.lean @@ -9,6 +9,9 @@ import geometry.euclidean.angle.oriented.basic /-! # Rotations by oriented angles. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines rotations by oriented angles in real inner product spaces. ## Main definitions diff --git a/src/geometry/euclidean/circumcenter.lean b/src/geometry/euclidean/circumcenter.lean index 41736c1778e66..a65c0db65410b 100644 --- a/src/geometry/euclidean/circumcenter.lean +++ b/src/geometry/euclidean/circumcenter.lean @@ -10,6 +10,9 @@ import tactic.derive_fintype /-! # Circumcenter and circumradius +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves some lemmas on points equidistant from a set of points, and defines the circumradius and circumcenter of a simplex. There are also some definitions for use in calculations where it is diff --git a/src/geometry/euclidean/sphere/ptolemy.lean b/src/geometry/euclidean/sphere/ptolemy.lean index aa5d6eaa4fec3..bb988789b7787 100644 --- a/src/geometry/euclidean/sphere/ptolemy.lean +++ b/src/geometry/euclidean/sphere/ptolemy.lean @@ -9,6 +9,9 @@ import geometry.euclidean.triangle /-! # Ptolemy's theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves Ptolemy's theorem on the lengths of the diagonals and sides of a cyclic quadrilateral. diff --git a/src/geometry/euclidean/triangle.lean b/src/geometry/euclidean/triangle.lean index 7e8e292c870a9..315ab548f31e4 100644 --- a/src/geometry/euclidean/triangle.lean +++ b/src/geometry/euclidean/triangle.lean @@ -10,6 +10,9 @@ import tactic.interval_cases /-! # Triangles +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves basic geometrical results about distances and angles in (possibly degenerate) triangles in real inner product spaces and Euclidean affine spaces. More specialized results, and results diff --git a/src/measure_theory/function/conditional_expectation/unique.lean b/src/measure_theory/function/conditional_expectation/unique.lean index 7d4488dcb41de..70986754eff09 100644 --- a/src/measure_theory/function/conditional_expectation/unique.lean +++ b/src/measure_theory/function/conditional_expectation/unique.lean @@ -9,6 +9,9 @@ import measure_theory.function.conditional_expectation.ae_measurable /-! # Uniqueness of the conditional expectation +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Two Lp functions `f, g` which are almost everywhere strongly measurable with respect to a σ-algebra `m` and verify `∫ x in s, f x ∂μ = ∫ x in s, g x ∂μ` for all `m`-measurable sets `s` are equal almost everywhere. This proves the uniqueness of the conditional expectation, which is not yet diff --git a/src/order/category/BddDistLat.lean b/src/order/category/BddDistLat.lean index 7523cd5d45659..1741ae9c5256a 100644 --- a/src/order/category/BddDistLat.lean +++ b/src/order/category/BddDistLat.lean @@ -9,6 +9,9 @@ import order.category.DistLat /-! # The category of bounded distributive lattices +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This defines `BddDistLat`, the category of bounded distributive lattices. Note that this category is sometimes called [`DistLat`](https://ncatlab.org/nlab/show/DistLat) when diff --git a/src/order/category/BddLat.lean b/src/order/category/BddLat.lean index 9ba37c0bf2829..99d624ae156e3 100644 --- a/src/order/category/BddLat.lean +++ b/src/order/category/BddLat.lean @@ -11,6 +11,9 @@ import order.category.Semilat /-! # The category of bounded lattices +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines `BddLat`, the category of bounded lattices. In literature, this is sometimes called `Lat`, the category of lattices, because being a lattice is diff --git a/src/order/category/CompleteLat.lean b/src/order/category/CompleteLat.lean index 1649a5cc2e6fc..1ff48a15c8536 100644 --- a/src/order/category/CompleteLat.lean +++ b/src/order/category/CompleteLat.lean @@ -9,6 +9,9 @@ import order.hom.complete_lattice /-! # The category of complete lattices +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines `CompleteLat`, the category of complete lattices. -/ diff --git a/src/order/category/HeytAlg.lean b/src/order/category/HeytAlg.lean index efea6f68b1608..bc4663ab9a5ea 100644 --- a/src/order/category/HeytAlg.lean +++ b/src/order/category/HeytAlg.lean @@ -9,6 +9,9 @@ import order.heyting.hom /-! # The category of Heyting algebras +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines `HeytAlg`, the category of Heyting algebras. -/ diff --git a/src/probability/kernel/measurable_integral.lean b/src/probability/kernel/measurable_integral.lean index 94b0146398a24..540b0aa13055e 100644 --- a/src/probability/kernel/measurable_integral.lean +++ b/src/probability/kernel/measurable_integral.lean @@ -8,6 +8,9 @@ import probability.kernel.basic /-! # Measurability of the integral against a kernel +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The Lebesgue integral of a measurable function against a kernel is measurable. The Bochner integral is strongly measurable. diff --git a/src/probability/kernel/with_density.lean b/src/probability/kernel/with_density.lean index 2b9b2724412d4..8a50da8a2afca 100644 --- a/src/probability/kernel/with_density.lean +++ b/src/probability/kernel/with_density.lean @@ -9,6 +9,9 @@ import measure_theory.integral.set_integral /-! # With Density +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + For an s-finite kernel `κ : kernel α β` and a function `f : α → β → ℝ≥0∞` which is finite everywhere, we define `with_density κ f` as the kernel `a ↦ (κ a).with_density (f a)`. This is an s-finite kernel. diff --git a/src/ring_theory/ring_hom/finite.lean b/src/ring_theory/ring_hom/finite.lean index 3959d106f125f..5b351cd00a69c 100644 --- a/src/ring_theory/ring_hom/finite.lean +++ b/src/ring_theory/ring_hom/finite.lean @@ -9,6 +9,9 @@ import ring_theory.ring_hom_properties # The meta properties of finite ring homomorphisms. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + -/ namespace ring_hom diff --git a/src/topology/continuous_function/ideals.lean b/src/topology/continuous_function/ideals.lean index c457752b83dbf..7341809a0f644 100644 --- a/src/topology/continuous_function/ideals.lean +++ b/src/topology/continuous_function/ideals.lean @@ -14,6 +14,9 @@ import topology.algebra.module.character_space /-! # Ideals of continuous functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + For a topological semiring `R` and a topological space `X` there is a Galois connection between `ideal C(X, R)` and `set X` given by sending each `I : ideal C(X, R)` to `{x : X | ∀ f ∈ I, f x = 0}ᶜ` and mapping `s : set X` to the ideal with carrier diff --git a/src/topology/gluing.lean b/src/topology/gluing.lean index 6186f6da75bf7..525390c4c3be2 100644 --- a/src/topology/gluing.lean +++ b/src/topology/gluing.lean @@ -11,6 +11,9 @@ import topology.category.Top.opens /-! # Gluing Topological spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given a family of gluing data (see `category_theory/glue_data`), we can then glue them together. The construction should be "sealed" and considered as a black box, while only using the API From e3f4be1fcb5376c4948d7f095bec45350bfb9d1a Mon Sep 17 00:00:00 2001 From: Riccardo Brasca Date: Thu, 15 Jun 2023 06:31:45 +0000 Subject: [PATCH 1189/1291] chore(field_theory/splitting_field): refactor `splitting_field` (#19178) We refactor the definition of `splitting_field`. The main motivation is to backport [!4#4891](https://github.com/leanprover-community/mathlib4/pull/4891) since the is seems very problematic to port the current design. Zulip discussion relevant to this PR: https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/!4.234891.20.20FieldTheory.2ESplittingField --- src/field_theory/abel_ruffini.lean | 4 +- src/field_theory/finite/galois_field.lean | 12 +- src/field_theory/polynomial_galois_group.lean | 25 +- .../splitting_field/construction.lean | 420 ++++++------------ src/number_theory/cyclotomic/basic.lean | 31 +- src/number_theory/cyclotomic/gal.lean | 10 +- .../legendre_symbol/add_character.lean | 37 +- .../legendre_symbol/gauss_sum.lean | 24 +- src/ring_theory/polynomial/gauss_lemma.lean | 3 +- 9 files changed, 228 insertions(+), 338 deletions(-) diff --git a/src/field_theory/abel_ruffini.lean b/src/field_theory/abel_ruffini.lean index 3f8412f818bab..aa5760145959a 100644 --- a/src/field_theory/abel_ruffini.lean +++ b/src/field_theory/abel_ruffini.lean @@ -344,8 +344,8 @@ begin exact minpoly.aeval F γ, end (minpoly.monic (is_integral γ)), rw [P, key], - exact gal_is_solvable_of_splits ⟨normal.splits (splitting_field.normal _) _⟩ - (gal_mul_is_solvable hα hβ), + refine gal_is_solvable_of_splits ⟨_⟩ (gal_mul_is_solvable hα hβ), + exact normal.splits (splitting_field.normal _) (f ⟨γ, hγ⟩), end /-- An auxiliary induction lemma, which is generalized by `solvable_by_rad.is_solvable`. -/ diff --git a/src/field_theory/finite/galois_field.lean b/src/field_theory/finite/galois_field.lean index 967ee213f8185..2199496911643 100644 --- a/src/field_theory/finite/galois_field.lean +++ b/src/field_theory/finite/galois_field.lean @@ -152,13 +152,15 @@ begin rw [splits_iff_card_roots, h1, ←finset.card_def, finset.card_univ, h2, zmod.card], end +local attribute [-instance] splitting_field.algebra' + /-- A Galois field with exponent 1 is equivalent to `zmod` -/ def equiv_zmod_p : galois_field p 1 ≃ₐ[zmod p] (zmod p) := -have h : (X ^ p ^ 1 : (zmod p)[X]) = X ^ (fintype.card (zmod p)), - by rw [pow_one, zmod.card p], -have inst : is_splitting_field (zmod p) (zmod p) (X ^ p ^ 1 - X), - by { rw h, apply_instance }, -by exactI (is_splitting_field.alg_equiv (zmod p) (X ^ (p ^ 1) - X : (zmod p)[X])).symm +let h : (X ^ p ^ 1 : (zmod p)[X]) = X ^ (fintype.card (zmod p)) := + by rw [pow_one, zmod.card p] in +let inst : is_splitting_field (zmod p) (zmod p) (X ^ p ^ 1 - X) := + by { rw h, apply_instance } in +(@is_splitting_field.alg_equiv _ (zmod p) _ _ _ (X ^ (p ^ 1) - X : (zmod p)[X]) inst).symm variables {K : Type*} [field K] [fintype K] [algebra (zmod p) K] diff --git a/src/field_theory/polynomial_galois_group.lean b/src/field_theory/polynomial_galois_group.lean index 76a94be444e84..d25a4855e0602 100644 --- a/src/field_theory/polynomial_galois_group.lean +++ b/src/field_theory/polynomial_galois_group.lean @@ -139,8 +139,7 @@ begin have hy := subtype.mem y, simp only [root_set, finset.mem_coe, multiset.mem_to_finset, key, multiset.mem_map] at hy, rcases hy with ⟨x, hx1, hx2⟩, - classical, - exact ⟨⟨x, multiset.mem_to_finset.mpr hx1⟩, subtype.ext hx2⟩ } + exact ⟨⟨x, (@multiset.mem_to_finset _ (classical.dec_eq _) _ _).mpr hx1⟩, subtype.ext hx2⟩ } end /-- The bijection between `root_set p p.splitting_field` and `root_set p E`. -/ @@ -237,14 +236,16 @@ begin { haveI : fact (p.splits (algebra_map F (p * q).splitting_field)) := ⟨splits_of_splits_of_dvd _ hpq (splitting_field.splits (p * q)) (dvd_mul_right p q)⟩, have key : x = algebra_map (p.splitting_field) (p * q).splitting_field - ((roots_equiv_roots p _).inv_fun ⟨x, multiset.mem_to_finset.mpr h⟩) := + ((roots_equiv_roots p _).inv_fun ⟨x, + (@multiset.mem_to_finset _ (classical.dec_eq _) _ _).mpr h⟩) := subtype.ext_iff.mp (equiv.apply_symm_apply (roots_equiv_roots p _) ⟨x, _⟩).symm, rw [key, ←alg_equiv.restrict_normal_commutes, ←alg_equiv.restrict_normal_commutes], exact congr_arg _ (alg_equiv.ext_iff.mp hfg.1 _) }, { haveI : fact (q.splits (algebra_map F (p * q).splitting_field)) := ⟨splits_of_splits_of_dvd _ hpq (splitting_field.splits (p * q)) (dvd_mul_left q p)⟩, have key : x = algebra_map (q.splitting_field) (p * q).splitting_field - ((roots_equiv_roots q _).inv_fun ⟨x, multiset.mem_to_finset.mpr h⟩) := + ((roots_equiv_roots q _).inv_fun ⟨x, + (@multiset.mem_to_finset _ (classical.dec_eq _) _ _).mpr h⟩) := subtype.ext_iff.mp (equiv.apply_symm_apply (roots_equiv_roots q _) ⟨x, _⟩).symm, rw [key, ←alg_equiv.restrict_normal_commutes, ←alg_equiv.restrict_normal_commutes], exact congr_arg _ (alg_equiv.ext_iff.mp hfg.2 _) }, @@ -257,12 +258,12 @@ lemma mul_splits_in_splitting_field_of_mul {p₁ q₁ p₂ q₂ : F[X]} (p₁ * p₂).splits (algebra_map F (q₁ * q₂).splitting_field) := begin apply splits_mul, - { rw ← (splitting_field.lift q₁ (splits_of_splits_of_dvd _ - (mul_ne_zero hq₁ hq₂) (splitting_field.splits _) (dvd_mul_right q₁ q₂))).comp_algebra_map, - exact splits_comp_of_splits _ _ h₁, }, - { rw ← (splitting_field.lift q₂ (splits_of_splits_of_dvd _ - (mul_ne_zero hq₁ hq₂) (splitting_field.splits _) (dvd_mul_left q₂ q₁))).comp_algebra_map, - exact splits_comp_of_splits _ _ h₂, }, + { rw ← (splitting_field.lift q₁ (splits_of_splits_of_dvd (algebra_map F (q₁ * q₂).splitting_field) + (mul_ne_zero hq₁ hq₂) (splitting_field.splits _) (dvd_mul_right q₁ q₂))).comp_algebra_map, + exact splits_comp_of_splits _ _ h₁ }, + { rw ← (splitting_field.lift q₂ (splits_of_splits_of_dvd (algebra_map F (q₁ * q₂).splitting_field) + (mul_ne_zero hq₁ hq₂) (splitting_field.splits _) (dvd_mul_left q₂ q₁))).comp_algebra_map, + exact splits_comp_of_splits _ _ h₂ }, end /-- `p` splits in the splitting field of `p ∘ q`, for `q` non-constant. -/ @@ -304,7 +305,9 @@ end /-- `polynomial.gal.restrict` for the composition of polynomials. -/ def restrict_comp (hq : q.nat_degree ≠ 0) : (p.comp q).gal →* p.gal := -@restrict F _ p _ _ _ ⟨splits_in_splitting_field_of_comp p q hq⟩ +let h : fact (splits (algebra_map F (p.comp q).splitting_field) p) := + ⟨splits_in_splitting_field_of_comp p q hq⟩ in +@restrict F _ p _ _ _ h lemma restrict_comp_surjective (hq : q.nat_degree ≠ 0) : function.surjective (restrict_comp p q hq) := diff --git a/src/field_theory/splitting_field/construction.lean b/src/field_theory/splitting_field/construction.lean index 6236de333f50c..6c9a641450ca0 100644 --- a/src/field_theory/splitting_field/construction.lean +++ b/src/field_theory/splitting_field/construction.lean @@ -19,6 +19,12 @@ In this file we prove the existence and uniqueness of splitting fields. * `polynomial.is_splitting_field.alg_equiv`: Every splitting field of a polynomial `f` is isomorphic to `splitting_field f` and thus, being a splitting field is unique up to isomorphism. +## Implementation details +We construct a `splitting_field_aux` without worrying about whether the instances satisfy nice +definitional equalities. Then the actual `splitting_field` is defined to be a quotient of a +`mv_polynomial` ring by the kernel of the obvious map into `splitting_field_aux`. Because the +actual `splitting_field` will be a quotient of a `mv_polynomial`, it has nice instances on it. + -/ noncomputable theory @@ -86,248 +92,38 @@ by rw [nat_degree_remove_factor, hfn, n.add_sub_cancel] /-- Auxiliary construction to a splitting field of a polynomial, which removes `n` (arbitrarily-chosen) factors. -Uses recursion on the degree. For better definitional behaviour, structures -including `splitting_field_aux` (such as instances) should be defined using -this recursion in each field, rather than defining the whole tuple through -recursion. --/ -def splitting_field_aux (n : ℕ) : Π {K : Type u} [field K], by exactI Π (f : K[X]), Type u := -nat.rec_on n (λ K _ _, K) $ λ n ih K _ f, by exactI -ih f.remove_factor - -namespace splitting_field_aux - -theorem succ (n : ℕ) (f : K[X]) : - splitting_field_aux (n+1) f = splitting_field_aux n f.remove_factor := rfl - -section lift_instances +It constructs the type, proves that is a field and algebra over the base field. -/-! ### Instances on `splitting_field_aux` - -In order to avoid diamond issues, we have to be careful to define each data field of algebraic -instances on `splitting_field_aux` by recursion, rather than defining the whole structure by -recursion at once. +Uses recursion on the degree. +-/ +def splitting_field_aux_aux (n : ℕ) : Π {K : Type u} [field K], by exactI Π (f : K[X]), + Σ (L : Type u) (inst : field L), by exactI algebra K L := +nat.rec_on n (λ K inst f, ⟨K, inst, infer_instance⟩) (λ m ih K inst f, + let L := ih (@remove_factor K inst f) in let h₁ := L.2.1 in let h₂ := L.2.2 in + ⟨L.1, L.2.1, by + { exactI (ring_hom.comp (algebra_map _ _) (adjoin_root.of f.factor)).to_algebra }⟩) -The important detail is that the `smul` instances can be lifted _before_ we create the algebraic -instances; if we do not do this, this creates a mutual dependence on both on each other, and it -is impossible to untangle this mess. In order to do this, we need that these `smul` instances -are distributive in the sense of `distrib_smul`, which is weak enough to be guaranteed at this -stage. This is sufficient to lift an action to `adjoin_root f` (remember that this is a quotient, -so these conditions are equivalent to well-definedness), which is all we need. +/-- Auxiliary construction to a splitting field of a polynomial, which removes +`n` (arbitrarily-chosen) factors. It is the type constructed in `splitting_field_aux_aux`. -/ +def splitting_field_aux (n : ℕ) {K : Type u} [field K] (f : K[X]) : Type u := + (splitting_field_aux_aux n f).1 -/-- Splitting fields have a zero. -/ -protected def zero (n : ℕ) : Π {K : Type u} [field K], by exactI Π {f : K[X]}, - splitting_field_aux n f := -nat.rec_on n (λ K fK f, by exactI @has_zero.zero K _) (λ n ih K fK f, ih) - -/-- Splitting fields have an addition. -/ -protected def add (n : ℕ) : Π {K : Type u} [field K], by exactI Π {f : K[X]}, - splitting_field_aux n f → splitting_field_aux n f → splitting_field_aux n f := -nat.rec_on n (λ K fK f, by exactI @has_add.add K _) (λ n ih K fK f, ih) - -/-- Splitting fields inherit scalar multiplication. -/ -protected def smul (n : ℕ) : Π (α : Type*) {K : Type u} [field K], - by exactI Π [distrib_smul α K], - by exactI Π [is_scalar_tower α K K] {f : K[X]}, - α → splitting_field_aux n f → splitting_field_aux n f := -nat.rec_on n - (λ α K fK ds ist f, by exactI @has_smul.smul _ K _) - (λ n ih α K fK ds ist f, by exactI ih α) - -instance has_smul (α : Type*) (n : ℕ) {K : Type u} [field K] [distrib_smul α K] - [is_scalar_tower α K K] {f : K[X]} : - has_smul α (splitting_field_aux n f) := -⟨splitting_field_aux.smul n α⟩ - -instance is_scalar_tower (n : ℕ) : Π (R₁ R₂ : Type*) {K : Type u} - [has_smul R₁ R₂] [field K], - by exactI Π [distrib_smul R₂ K] [distrib_smul R₁ K], - by exactI Π [is_scalar_tower R₁ K K] [is_scalar_tower R₂ K K] [is_scalar_tower R₁ R₂ K] - {f : K[X]}, by exactI is_scalar_tower R₁ R₂ (splitting_field_aux n f) := -nat.rec_on n (λ R₁ R₂ K _ _ hs₂ hs₁ _ _ h f, -begin - rcases hs₁ with @⟨@⟨⟨hs₁⟩,_⟩,_⟩, - rcases hs₂ with @⟨@⟨⟨hs₂⟩,_⟩,_⟩, - exact h, -end) $ λ n ih R₁ R₂ K _ _ _ _ _ _ _ f, by exactI ih R₁ R₂ - -/-- Splitting fields have a negation. -/ -protected def neg (n : ℕ) : Π {K : Type u} [field K], by exactI Π {f : K[X]}, - splitting_field_aux n f → splitting_field_aux n f := -nat.rec_on n (λ K fK f, by exactI @has_neg.neg K _) (λ n ih K fK f, ih) - -/-- Splitting fields have subtraction. -/ -protected def sub (n : ℕ) : Π {K : Type u} [field K], by exactI Π {f : K[X]}, - splitting_field_aux n f → splitting_field_aux n f → splitting_field_aux n f := -nat.rec_on n (λ K fK f, by exactI @has_sub.sub K _) (λ n ih K fK f, ih) - -/-- Splitting fields have a one. -/ -protected def one (n : ℕ) : Π {K : Type u} [field K], by exactI Π {f : K[X]}, - splitting_field_aux n f := -nat.rec_on n (λ K fK f, by exactI @has_one.one K _) (λ n ih K fK f, ih) - -/-- Splitting fields have a multiplication. -/ -protected def mul (n : ℕ) : Π {K : Type u} [field K], by exactI Π {f : K[X]}, - splitting_field_aux n f → splitting_field_aux n f → splitting_field_aux n f := -nat.rec_on n (λ K fK f, by exactI @has_mul.mul K _) (λ n ih K fK f, ih) - -/-- Splitting fields have a power operation. -/ -protected def npow (n : ℕ) : Π {K : Type u} [field K], by exactI Π {f : K[X]}, - ℕ → splitting_field_aux n f → splitting_field_aux n f := -nat.rec_on n (λ K fK f n x, by exactI @has_pow.pow K _ _ x n) (λ n ih K fK f, ih) - -/-- Splitting fields have an injection from the base field. -/ -protected def mk (n : ℕ) : Π {K : Type u} [field K], - by exactI Π {f : K[X]}, K → splitting_field_aux n f := -nat.rec_on n (λ K fK f, id) (λ n ih K fK f, by exactI ih ∘ coe) --- note that `coe` goes from `K → adjoin_root f`, and then `ih` lifts to the full splitting field --- (as we generalise over all fields in this recursion!) - -/-- Splitting fields have an inverse. -/ -protected def inv (n : ℕ) : Π {K : Type u} [field K], by exactI Π {f : K[X]}, - splitting_field_aux n f → splitting_field_aux n f := -nat.rec_on n (λ K fK f, by exactI @has_inv.inv K _) (λ n ih K fK f, ih) - -/-- Splitting fields have a division. -/ -protected def div (n : ℕ) : Π {K : Type u} [field K], by exactI Π {f : K[X]}, - splitting_field_aux n f → splitting_field_aux n f → splitting_field_aux n f := -nat.rec_on n (λ K fK f, by exactI @has_div.div K _) (λ n ih K fK f, ih) - -/-- Splitting fields have powers by integers. -/ -protected def zpow (n : ℕ) : Π {K : Type u} [field K], by exactI Π {f : K[X]}, - ℤ → splitting_field_aux n f → splitting_field_aux n f := -nat.rec_on n (λ K fK f n x, by exactI @has_pow.pow K _ _ x n) (λ n ih K fK f, ih) - --- I'm not sure why these two lemmas break, but inlining them seems to not work. -private lemma nsmul_zero (n : ℕ) : Π {K : Type u} [field K], by exactI Π {f : K[X]} - (x : splitting_field_aux n f), (0 : ℕ) • x = splitting_field_aux.zero n := -nat.rec_on n (λ K fK f, by exactI zero_nsmul) (λ n ih K fK f, by exactI ih) - -private lemma nsmul_succ (n : ℕ) : Π {K : Type u} [field K], by exactI Π {f : K[X]} - (k : ℕ) (x : splitting_field_aux n f), - (k + 1) • x = splitting_field_aux.add n x (k • x) := -nat.rec_on n (λ K fK f n x, by exactI succ_nsmul x n) (λ n ih K fK f, by exactI ih) - -instance field (n : ℕ) {K : Type u} [field K] {f : K[X]} : - field (splitting_field_aux n f) := -begin - refine - { zero := splitting_field_aux.zero n, - one := splitting_field_aux.one n, - add := splitting_field_aux.add n, - zero_add := have h : _ := @zero_add, _, - add_zero := have h : _ := @add_zero, _, - add_assoc := have h : _ := @add_assoc, _, - add_comm := have h : _ := @add_comm, _, - neg := splitting_field_aux.neg n, - add_left_neg := have h : _ := @add_left_neg, _, - sub := splitting_field_aux.sub n, - sub_eq_add_neg := have h : _ := @sub_eq_add_neg, _, - mul := splitting_field_aux.mul n, - one_mul := have h : _ := @one_mul, _, - mul_one := have h : _ := @mul_one, _, - mul_assoc := have h : _ := @mul_assoc, _, - left_distrib := have h : _ := @left_distrib, _, - right_distrib := have h : _ := @right_distrib, _, - mul_comm := have h : _ := @mul_comm, _, - inv := splitting_field_aux.inv n, - inv_zero := have h : _ := @inv_zero, _, - div := splitting_field_aux.div n, - div_eq_mul_inv := have h : _ := @div_eq_mul_inv, _, - mul_inv_cancel := have h : _ := @mul_inv_cancel, _, - exists_pair_ne := have h : _ := @exists_pair_ne, _, - nat_cast := splitting_field_aux.mk n ∘ (coe : ℕ → K), - nat_cast_zero := have h : _ := @comm_ring.nat_cast_zero, _, - nat_cast_succ := have h : _ := @comm_ring.nat_cast_succ, _, - nsmul := (•), - nsmul_zero' := nsmul_zero n, - nsmul_succ' := nsmul_succ n, - int_cast := splitting_field_aux.mk n ∘ (coe : ℤ → K), - int_cast_of_nat := have h : _ := @comm_ring.int_cast_of_nat, _, - int_cast_neg_succ_of_nat := have h : _ := @comm_ring.int_cast_neg_succ_of_nat, _, - zsmul := (•), - zsmul_zero' := have h : _ := @subtraction_comm_monoid.zsmul_zero', _, - zsmul_succ' := have h : _ := @subtraction_comm_monoid.zsmul_succ', _, - zsmul_neg' := have h : _ := @subtraction_comm_monoid.zsmul_neg', _, - rat_cast := splitting_field_aux.mk n ∘ (coe : ℚ → K), - rat_cast_mk := have h : _ := @field.rat_cast_mk, _, - qsmul := (•), - qsmul_eq_mul' := have h : _ := @field.qsmul_eq_mul', _, - npow := splitting_field_aux.npow n, - npow_zero' := have h : _ := @comm_ring.npow_zero', _, - npow_succ' := have h : _ := @comm_ring.npow_succ', _, - zpow := splitting_field_aux.zpow n, - zpow_zero' := have h : _ := @field.zpow_zero', _, - zpow_succ' := have h : _ := @field.zpow_succ', _, - zpow_neg' := have h : _ := @field.zpow_neg', _}, - all_goals { - unfreezingI { induction n with n ih generalizing K }, - { apply @h K }, - { exact @ih _ _ _ } }, -end +instance splitting_field_aux.field (n : ℕ) {K : Type u} [field K] (f : K[X]) : + field (splitting_field_aux n f) := + (splitting_field_aux_aux n f).2.1 -instance inhabited {n : ℕ} {f : K[X]} : - inhabited (splitting_field_aux n f) := ⟨37⟩ +instance (n : ℕ) {K : Type u} [field K] (f : K[X]) : inhabited (splitting_field_aux n f) := +⟨0⟩ -/-- The injection from the base field as a ring homomorphism. -/ -@[simps] protected def mk_hom (n : ℕ) {K : Type u} [field K] {f : K[X]} : - K →+* splitting_field_aux n f := -{ to_fun := splitting_field_aux.mk n, - map_one' := - begin - unfreezingI { induction n with k hk generalizing K }, - { simp [splitting_field_aux.mk] }, - exact hk - end, - map_mul' := - begin - unfreezingI { induction n with k hk generalizing K }, - { simp [splitting_field_aux.mk] }, - intros x y, - change (splitting_field_aux.mk k) ((adjoin_root.of f.factor) _) = _, - rw [map_mul], - exact hk _ _ - end, - map_zero' := - begin - unfreezingI { induction n with k hk generalizing K }, - { simp [splitting_field_aux.mk] }, - change (splitting_field_aux.mk k) ((adjoin_root.of f.factor) 0) = _, - rw [map_zero, hk], - end, - map_add' := - begin - unfreezingI { induction n with k hk generalizing K }, - { simp [splitting_field_aux.mk] }, - intros x y, - change (splitting_field_aux.mk k) ((adjoin_root.of f.factor) _) = _, - rw [map_add], - exact hk _ _ - end } - -instance algebra (n : ℕ) (R : Type*) {K : Type u} [comm_semiring R] [field K] - [algebra R K] {f : K[X]} : algebra R (splitting_field_aux n f) := -{ to_fun := (splitting_field_aux.mk_hom n).comp (algebra_map R K), - smul := (•), - smul_def' := - begin - unfreezingI { induction n with k hk generalizing K }, - { exact algebra.smul_def }, - exact hk - end, - commutes' := λ a b, mul_comm _ _, - .. (splitting_field_aux.mk_hom n).comp (algebra_map R K) } - -/-- Because `splitting_field_aux` is defined by recursion, we have to make sure all instances -on `splitting_field_aux` are defined by recursion within the fields. Otherwise, there will be -instance diamonds such as the following: -/ -example (n : ℕ) {K : Type u} [field K] {f : K[X]} : - (add_comm_monoid.nat_module : module ℕ (splitting_field_aux n f)) = - @algebra.to_module _ _ _ _ (splitting_field_aux.algebra n _) := -rfl +instance splitting_field_aux.algebra (n : ℕ) {K : Type u} [field K] (f : K[X]) : + algebra K (splitting_field_aux n f) := + (splitting_field_aux_aux n f).2.2 -end lift_instances +namespace splitting_field_aux + +theorem succ (n : ℕ) (f : K[X]) : + splitting_field_aux (n+1) f = splitting_field_aux n f.remove_factor := rfl instance algebra''' {n : ℕ} {f : K[X]} : algebra (adjoin_root f.factor) @@ -340,29 +136,19 @@ splitting_field_aux.algebra''' instance algebra'' {n : ℕ} {f : K[X]} : algebra K (splitting_field_aux n f.remove_factor) := -splitting_field_aux.algebra n K +ring_hom.to_algebra (ring_hom.comp (algebra_map _ _) (adjoin_root.of f.factor)) instance scalar_tower' {n : ℕ} {f : K[X]} : is_scalar_tower K (adjoin_root f.factor) (splitting_field_aux n f.remove_factor) := -begin - -- finding this instance ourselves makes things faster - haveI : is_scalar_tower K (adjoin_root f.factor) (adjoin_root f.factor) := - is_scalar_tower.right, - exact - splitting_field_aux.is_scalar_tower n K (adjoin_root f.factor), -end - -instance scalar_tower {n : ℕ} {f : K[X]} : - is_scalar_tower K (adjoin_root f.factor) (splitting_field_aux (n + 1) f) := -splitting_field_aux.scalar_tower' +is_scalar_tower.of_algebra_map_eq (λ x, rfl) theorem algebra_map_succ (n : ℕ) (f : K[X]) : by exact algebra_map K (splitting_field_aux (n+1) f) = (algebra_map (adjoin_root f.factor) (splitting_field_aux n f.remove_factor)).comp (adjoin_root.of f.factor) := -is_scalar_tower.algebra_map_eq _ _ _ +rfl protected theorem splits (n : ℕ) : ∀ {K : Type u} [field K], by exactI ∀ (f : K[X]) (hfn : f.nat_degree = n), @@ -373,24 +159,6 @@ by { resetI, rw [← splits_id_iff_splits, algebra_map_succ, ← map_map, splits ← X_sub_C_mul_remove_factor f (λ h, by { rw h at hf, cases hf })], exact splits_mul _ (splits_X_sub_C _) (ih _ (nat_degree_remove_factor' hf)) } -theorem exists_lift (n : ℕ) : ∀ {K : Type u} [field K], by exactI - ∀ (f : K[X]) (hfn : f.nat_degree = n) {L : Type*} [field L], by exactI - ∀ (j : K →+* L) (hf : splits j f), ∃ k : splitting_field_aux n f →+* L, - k.comp (algebra_map _ _) = j := -nat.rec_on n (λ K _ _ _ L _ j _, by exactI ⟨j, j.comp_id⟩) $ λ n ih K _ f hf L _ j hj, by exactI -have hndf : f.nat_degree ≠ 0, by { intro h, rw h at hf, cases hf }, -have hfn0 : f ≠ 0, by { intro h, rw h at hndf, exact hndf rfl }, -let ⟨r, hr⟩ := exists_root_of_splits _ (splits_of_splits_of_dvd j hfn0 hj - (factor_dvd_of_nat_degree_ne_zero hndf)) - (mt is_unit_iff_degree_eq_zero.2 f.irreducible_factor.1) in -have hmf0 : map (adjoin_root.of f.factor) f ≠ 0, from map_ne_zero hfn0, -have hsf : splits (adjoin_root.lift j r hr) f.remove_factor, -by { rw ← X_sub_C_mul_remove_factor _ hndf at hmf0, refine (splits_of_splits_mul _ hmf0 _).2, - rwa [X_sub_C_mul_remove_factor _ hndf, ← splits_id_iff_splits, map_map, adjoin_root.lift_comp_of, - splits_id_iff_splits] }, -let ⟨k, hk⟩ := ih f.remove_factor (nat_degree_remove_factor' hf) (adjoin_root.lift j r hr) hsf in -⟨k, by rw [algebra_map_succ, ← ring_hom.comp_assoc, hk, adjoin_root.lift_comp_of]⟩ - theorem adjoin_root_set (n : ℕ) : ∀ {K : Type u} [field K], by exactI ∀ (f : K[X]) (hfn : f.nat_degree = n), algebra.adjoin K (f.root_set (splitting_field_aux n f)) = ⊤ := @@ -414,26 +182,108 @@ begin ih _ (nat_degree_remove_factor' hfn), subalgebra.restrict_scalars_top] end +instance (f : K[X]) : is_splitting_field K (splitting_field_aux f.nat_degree f) f := + ⟨splitting_field_aux.splits _ _ rfl, splitting_field_aux.adjoin_root_set _ _ rfl⟩ + +/-- The natural map from `mv_polynomial (f.root_set (splitting_field_aux f.nat_degree f))` +to `splitting_field_aux f.nat_degree f` sendind a variable to the corresponding root. -/ +def of_mv_polynomial (f : K[X]) : + mv_polynomial (f.root_set (splitting_field_aux f.nat_degree f)) K →ₐ[K] + splitting_field_aux f.nat_degree f := + mv_polynomial.aeval (λ i, i.1) + +theorem of_mv_polynomial_surjective (f : K[X]) : function.surjective (of_mv_polynomial f) := +begin + suffices : alg_hom.range (of_mv_polynomial f) = ⊤, + { rw [← set.range_iff_surjective]; rwa [set_like.ext'_iff] at this }, + rw [of_mv_polynomial, ← algebra.adjoin_range_eq_range_aeval K, eq_top_iff, + ← adjoin_root_set _ _ rfl], + exact algebra.adjoin_le (λ α hα, algebra.subset_adjoin ⟨⟨α, hα⟩, rfl⟩) +end + +/-- The algebra isomorphism between the quotient of +`mv_polynomial (f.root_set (splitting_field_aux f.nat_degree f)) K` by the kernel of +`of_mv_polynomial f` and `splitting_field_aux f.nat_degree f`. It is used to transport all the +algebraic structures from the latter to `f.splitting_field`, that is defined as the former. -/ +def alg_equiv_quotient_mv_polynomial (f : K[X]) : + (mv_polynomial (f.root_set (splitting_field_aux f.nat_degree f)) K ⧸ + ring_hom.ker (of_mv_polynomial f).to_ring_hom) ≃ₐ[K] + splitting_field_aux f.nat_degree f := + (ideal.quotient_ker_alg_equiv_of_surjective (of_mv_polynomial_surjective f) : _) + end splitting_field_aux /-- A splitting field of a polynomial. -/ def splitting_field (f : K[X]) := -splitting_field_aux f.nat_degree f +mv_polynomial (f.root_set (splitting_field_aux f.nat_degree f)) K ⧸ + ring_hom.ker (splitting_field_aux.of_mv_polynomial f).to_ring_hom namespace splitting_field variables (f : K[X]) -instance : field (splitting_field f) := -splitting_field_aux.field _ +instance comm_ring : comm_ring (splitting_field f) := +ideal.quotient.comm_ring _ + +instance inhabited : inhabited (splitting_field f) := + ⟨37⟩ -instance inhabited : inhabited (splitting_field f) := ⟨37⟩ +instance {S : Type*} [distrib_smul S K] [is_scalar_tower S K K] : + has_smul S (splitting_field f) := + submodule.quotient.has_smul' _ -instance algebra' {R} [comm_semiring R] [algebra R K] : algebra R (splitting_field f) := -splitting_field_aux.algebra _ _ +instance algebra : algebra K (splitting_field f) := +ideal.quotient.algebra _ -instance : algebra K (splitting_field f) := -splitting_field_aux.algebra _ _ +instance algebra' {R : Type*} [comm_semiring R] [algebra R K] : algebra R (splitting_field f) := +ideal.quotient.algebra _ + +instance is_scalar_tower {R : Type*} [comm_semiring R] [algebra R K] : + is_scalar_tower R K (splitting_field f) := +ideal.quotient.is_scalar_tower _ _ _ + +/-- The algebra equivalence with `splitting_field_aux`, +which we will use to construct the field structure. -/ +def alg_equiv_splitting_field_aux (f : K[X]) : + splitting_field f ≃ₐ[K] splitting_field_aux f.nat_degree f := + splitting_field_aux.alg_equiv_quotient_mv_polynomial f + +instance : field (splitting_field f) := +let e := alg_equiv_splitting_field_aux f in +{ rat_cast := λ a, algebra_map K _ (a : K), + inv := λ a, e.symm (e a)⁻¹, + qsmul := (•), + qsmul_eq_mul' := λ a x, quotient.induction_on' x (λ p, congr_arg quotient.mk' + begin + ext, + simp only [mv_polynomial.algebra_map_eq, rat.smul_def, mv_polynomial.coeff_smul, + mv_polynomial.coeff_C_mul], + end), + rat_cast_mk := λ a b h1 h2, + begin + apply e.injective, + change e (algebra_map K _ _) = _, + simp only [map_rat_cast, map_nat_cast, map_mul, map_int_cast, alg_equiv.commutes], + change _ = e ↑a * e (e.symm (e b)⁻¹), + rw [alg_equiv.apply_symm_apply], + convert field.rat_cast_mk a b h1 h2, + all_goals { simp }, + end, + exists_pair_ne := ⟨e.symm 0, e.symm 1, λ w, zero_ne_one ((e.symm).injective w)⟩, + mul_inv_cancel := λ a w, + begin + apply e.injective, + rw [map_mul, map_one], + change e a * e (e.symm (e a)⁻¹) = 1, + rw [alg_equiv.apply_symm_apply, mul_inv_cancel], + exact (λ w', w (by simpa only [add_equiv_class.map_eq_zero_iff] using w')), + end, + inv_zero := + begin + change e.symm (e 0)⁻¹ = 0, + simp + end, + ..splitting_field.comm_ring f } instance [char_zero K] : char_zero (splitting_field f) := char_zero_of_injective_algebra_map ((algebra_map K _).injective) @@ -451,24 +301,23 @@ rfl example [char_zero K] : (splitting_field.algebra' f) = algebra_rat := rfl -example [char_zero K] : (splitting_field.algebra' f) = algebra_rat := -rfl - example {q : ℚ[X]} : algebra_int (splitting_field q) = splitting_field.algebra' q := rfl +instance _root_.polynomial.is_splitting_field.splitting_field (f : K[X]) : + is_splitting_field K (splitting_field f) f := + is_splitting_field.of_alg_equiv _ f (splitting_field_aux.alg_equiv_quotient_mv_polynomial f).symm + protected theorem splits : splits (algebra_map K (splitting_field f)) f := -splitting_field_aux.splits _ _ rfl +is_splitting_field.splits f.splitting_field f variables [algebra K L] (hb : splits (algebra_map K L) f) /-- Embeds the splitting field into any other field that splits the polynomial. -/ def lift : splitting_field f →ₐ[K] L := -{ commutes' := λ r, by { have := classical.some_spec (splitting_field_aux.exists_lift _ _ rfl _ hb), - exact ring_hom.ext_iff.1 this r }, - .. classical.some (splitting_field_aux.exists_lift _ _ _ _ hb) } +is_splitting_field.lift f.splitting_field f hb theorem adjoin_root_set : algebra.adjoin K (f.root_set (splitting_field f)) = ⊤ := -splitting_field_aux.adjoin_root_set _ _ rfl +polynomial.is_splitting_field.adjoin_root_set _ f end splitting_field @@ -479,12 +328,15 @@ namespace is_splitting_field variables (K L) [algebra K L] variables {K} -instance splitting_field (f : K[X]) : is_splitting_field K (splitting_field f) f := -⟨splitting_field.splits f, splitting_field.adjoin_root_set f⟩ instance (f : K[X]) : _root_.finite_dimensional K f.splitting_field := finite_dimensional f.splitting_field f +instance [fintype K] (f : K[X]) : fintype f.splitting_field := +finite_dimensional.fintype_of_fintype K _ + +instance (f : K[X]) : no_zero_smul_divisors K f.splitting_field := infer_instance + /-- Any splitting field is isomorphic to `splitting_field f`. -/ def alg_equiv (f : K[X]) [is_splitting_field K L f] : L ≃ₐ[K] splitting_field f := begin diff --git a/src/number_theory/cyclotomic/basic.lean b/src/number_theory/cyclotomic/basic.lean index da9fcb87c7960..b30c3a02e03b4 100644 --- a/src/number_theory/cyclotomic/basic.lean +++ b/src/number_theory/cyclotomic/basic.lean @@ -479,12 +479,15 @@ lemma splitting_field_X_pow_sub_one : is_splitting_field K L (X ^ (n : ℕ) - 1) exact λ _, X_pow_sub_C_ne_zero n.pos (1 : L) end } +local attribute [-instance] splitting_field.algebra' + /-- Any two `n`-th cyclotomic extensions are isomorphic. -/ def alg_equiv (L' : Type*) [field L'] [algebra K L'] [is_cyclotomic_extension {n} K L'] : L ≃ₐ[K] L' := -let _ := splitting_field_X_pow_sub_one n K L in let _ := splitting_field_X_pow_sub_one n K L' in - by exactI (is_splitting_field.alg_equiv L (X ^ (n : ℕ) - 1)).trans - (is_splitting_field.alg_equiv L' (X ^ (n : ℕ) - 1)).symm +let h₁ := splitting_field_X_pow_sub_one n K L in +let h₂ := splitting_field_X_pow_sub_one n K L' in +((@is_splitting_field.alg_equiv K L _ _ _ (X ^ (n : ℕ) - 1) h₁).trans + (@is_splitting_field.alg_equiv K L' _ _ _ (X ^ (n : ℕ) - 1) h₂).symm) localized "attribute [instance] is_cyclotomic_extension.splitting_field_X_pow_sub_one" in cyclotomic @@ -552,7 +555,7 @@ end cyclotomic_field section is_domain -variables [is_domain A] [algebra A K] [is_fraction_ring A K] +variables [algebra A K] [is_fraction_ring A K] section cyclotomic_ring @@ -565,14 +568,26 @@ splitting_field.algebra' (cyclotomic n K) /-- Ensure there are no diamonds when `A = ℤ`. -/ example : algebra_int (cyclotomic_field n ℚ) = cyclotomic_field.algebra_base _ _ _ := rfl +instance cyclotomic_field.algebra' {R : Type*} [comm_ring R] [algebra R K] : + algebra R (cyclotomic_field n K) := +splitting_field.algebra' (cyclotomic n K) + +instance {R : Type*} [comm_ring R] [algebra R K] : is_scalar_tower R K (cyclotomic_field n K) := +splitting_field.is_scalar_tower _ + instance cyclotomic_field.no_zero_smul_divisors : no_zero_smul_divisors A (cyclotomic_field n K) := -no_zero_smul_divisors.of_algebra_map_injective $ function.injective.comp -(no_zero_smul_divisors.algebra_map_injective _ _) $ is_fraction_ring.injective A K +begin + refine no_zero_smul_divisors.of_algebra_map_injective _, + rw is_scalar_tower.algebra_map_eq A K (cyclotomic_field n K), + exact function.injective.comp + (no_zero_smul_divisors.algebra_map_injective K (cyclotomic_field n K)) + (is_fraction_ring.injective A K), +end /-- If `A` is a domain with fraction field `K` and `n : ℕ+`, we define `cyclotomic_ring n A K` as the `A`-subalgebra of `cyclotomic_field n K` generated by the roots of `X ^ n - 1`. If `n` is nonzero in `A`, it has the instance `is_cyclotomic_extension {n} A (cyclotomic_ring n A K)`. -/ -@[derive [comm_ring, is_domain, inhabited]] +@[derive [comm_ring, is_domain, inhabited], nolint unused_arguments] def cyclotomic_ring : Type w := adjoin A { b : (cyclotomic_field n K) | b ^ (n : ℕ) = 1 } namespace cyclotomic_ring @@ -628,7 +643,7 @@ instance is_cyclotomic_extension [ne_zero ((n : ℕ) : A)] : { exact subalgebra.mul_mem _ hy hz }, end } -instance [ne_zero ((n : ℕ) : A)] : +instance [is_domain A] [ne_zero ((n : ℕ) : A)] : is_fraction_ring (cyclotomic_ring n A K) (cyclotomic_field n K) := { map_units := λ ⟨x, hx⟩, begin rw is_unit_iff_ne_zero, diff --git a/src/number_theory/cyclotomic/gal.lean b/src/number_theory/cyclotomic/gal.lean index e9df88c21c521..ee16e36e0619c 100644 --- a/src/number_theory/cyclotomic/gal.lean +++ b/src/number_theory/cyclotomic/gal.lean @@ -160,15 +160,17 @@ Asserts that the Galois group of `cyclotomic n K` is equivalent to `(zmod n)ˣ` if `cyclotomic n K` is irreducible in the base field. -/ noncomputable def gal_cyclotomic_equiv_units_zmod : (cyclotomic n K).gal ≃* (zmod n)ˣ := -(alg_equiv.aut_congr (is_splitting_field.alg_equiv _ _)).symm.trans -(is_cyclotomic_extension.aut_equiv_pow L h) +(alg_equiv.aut_congr + ((is_splitting_field.alg_equiv L _) : L ≃ₐ[K] (cyclotomic n K).splitting_field)).symm.trans + (is_cyclotomic_extension.aut_equiv_pow L h) /-- `is_cyclotomic_extension.aut_equiv_pow` repackaged in terms of `gal`. Asserts that the Galois group of `X ^ n - 1` is equivalent to `(zmod n)ˣ` if `cyclotomic n K` is irreducible in the base field. -/ noncomputable def gal_X_pow_equiv_units_zmod : (X ^ (n : ℕ) - 1).gal ≃* (zmod n)ˣ := -(alg_equiv.aut_congr (is_splitting_field.alg_equiv _ _)).symm.trans -(is_cyclotomic_extension.aut_equiv_pow L h) +(alg_equiv.aut_congr + ((is_splitting_field.alg_equiv L _) : L ≃ₐ[K] (X ^ (n : ℕ) - 1).splitting_field)).symm.trans + (is_cyclotomic_extension.aut_equiv_pow L h) end gal diff --git a/src/number_theory/legendre_symbol/add_character.lean b/src/number_theory/legendre_symbol/add_character.lean index 6e898d7f228b7..94407f97e1d82 100644 --- a/src/number_theory/legendre_symbol/add_character.lean +++ b/src/number_theory/legendre_symbol/add_character.lean @@ -224,14 +224,30 @@ begin rwa [mul_shift_apply, mul_inv_cancel_left₀ ha], end -/-- Structure for a primitive additive character on a finite ring `R` into a cyclotomic extension +/-- Definition for a primitive additive character on a finite ring `R` into a cyclotomic extension of a field `R'`. It records which cyclotomic extension it is, the character, and the fact that the character is primitive. -/ +-- Using `structure` gives a timeout, see +-- https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/mysterious.20finsupp.20related.20timeout/near/365719262 and +-- https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/mysterious.20finsupp.20related.20timeout +-- In Lean4, `set_option genInjectivity false in` may solve this issue. @[nolint has_nonempty_instance] -- can't prove that they always exist -structure primitive_add_char (R : Type u) [comm_ring R] [fintype R] (R' : Type v) [field R'] := -(n : ℕ+) -(char : add_char R (cyclotomic_field n R')) -(prim : is_primitive char) +def primitive_add_char (R : Type u) [comm_ring R] (R' : Type v) [field R'] := +Σ (n : ℕ+), (Σ' (char : add_char R (cyclotomic_field n R')), is_primitive char) + +/-- The first projection from `primitive_add_char`, giving the cyclotomic field. -/ +noncomputable! def primitive_add_char.n {R : Type u} [comm_ring R] {R' : Type v} + [field R'] : primitive_add_char R R' → ℕ+ := λ χ, χ.1 + +/-- The second projection from `primitive_add_char`, giving the character. -/ +noncomputable! def primitive_add_char.char {R : Type u} [comm_ring R] {R' : Type v} + [field R'] : Π (χ : primitive_add_char R R'), add_char R (cyclotomic_field χ.n R') := + λ χ, χ.2.1 + +/-- The third projection from `primitive_add_char`, showing that `χ.2` is primitive. -/ +lemma primitive_add_char.prim {R : Type u} [comm_ring R] {R' : Type v} + [field R'] : Π (χ : primitive_add_char R R'), is_primitive χ.char := + λ χ, χ.2.2 /-! ### Additive characters on `zmod n` @@ -311,10 +327,8 @@ def primitive_zmod_char (n : ℕ+) (F' : Type v) [field F'] (h : (n : F') ≠ 0) primitive_add_char (zmod n) F' := begin haveI : ne_zero ((n : ℕ) : F') := ⟨h⟩, - exact -{ n := n, - char := zmod_char n (is_cyclotomic_extension.zeta_pow n F' _), - prim := zmod_char_primitive_of_primitive_root n (is_cyclotomic_extension.zeta_spec n F' _) } + exact ⟨n, zmod_char n (is_cyclotomic_extension.zeta_pow n F' _), + zmod_char_primitive_of_primitive_root n (is_cyclotomic_extension.zeta_spec n F' _)⟩ end /-! @@ -348,10 +362,7 @@ begin rw one_mul at ha, exact ⟨a, λ hf, ha $ (ψ.prim.zmod_char_eq_one_iff pp $ algebra.trace (zmod p) F a).mp hf⟩, end, - exact -{ n := ψ.n, - char := ψ', - prim := hψ'.is_primitive }, + exact ⟨ψ.n, ψ', hψ'.is_primitive⟩ end /-! diff --git a/src/number_theory/legendre_symbol/gauss_sum.lean b/src/number_theory/legendre_symbol/gauss_sum.lean index 17b8fa9862a53..43f375c18d498 100644 --- a/src/number_theory/legendre_symbol/gauss_sum.lean +++ b/src/number_theory/legendre_symbol/gauss_sum.lean @@ -285,16 +285,22 @@ begin -- we now show that the Gauss sum of `χ` and `ψ₈` has the relevant property have hg : gauss_sum χ ψ₈.char ^ 2 = χ (-1) * fintype.card (zmod 8), - { rw [hχ, one_mul, card, gauss_sum], - convert ← congr_arg (^ 2) (fin.sum_univ_eight $ λ x, (χ₈ x : FF) * τ ^ x.val), + { have h := congr_arg (^ 2) (fin.sum_univ_eight $ λ x, (χ₈ x : FF) * τ ^ x.1), + have h₁ : (λ (i : fin 8), ↑(χ₈ i) * τ ^ i.val) = λ (a : zmod 8), χ a * ψ₈.char a, { ext, congr, apply pow_one }, - convert_to (0 + 1 * τ ^ 1 + 0 + (-1) * τ ^ 3 + 0 + (-1) * τ ^ 5 + 0 + 1 * τ ^ 7) ^ 2 = _, - { simp only [χ₈_apply, matrix.cons_val_zero, matrix.cons_val_one, matrix.head_cons, - matrix.cons_vec_bit0_eq_alt0, matrix.cons_vec_bit1_eq_alt1, matrix.cons_vec_append, - matrix.cons_vec_alt0, matrix.cons_vec_alt1, int.cast_zero, int.cast_one, int.cast_neg, - zero_mul], refl }, - convert_to 8 + (τ ^ 4 + 1) * (τ ^ 10 - 2 * τ ^ 8 - 2 * τ ^ 6 + 6 * τ ^ 4 + τ ^ 2 - 8) = _, - { ring }, { rw τ_spec, norm_num } }, + have h₂ : (0 + 1 * τ ^ 1 + 0 + (-1) * τ ^ 3 + 0 + (-1) * τ ^ 5 + 0 + 1 * τ ^ 7) ^ 2 = + 8 + (τ ^ 4 + 1) * (τ ^ 10 - 2 * τ ^ 8 - 2 * τ ^ 6 + 6 * τ ^ 4 + τ ^ 2 - 8) := by ring, + have h₃ : 8 + (τ ^ 4 + 1) * (τ ^ 10 - 2 * τ ^ 8 - 2 * τ ^ 6 + 6 * τ ^ 4 + τ ^ 2 - 8) = + ↑8 := by { rw τ_spec, norm_num }, + have h₄ : (0 + 1 * τ ^ 1 + 0 + (-1) * τ ^ 3 + 0 + (-1) * τ ^ 5 + 0 + 1 * τ ^ 7) ^ 2 = ↑8, + { rw [← h₃, ← h₂] }, + have h₅ : (λ (x : FF), x ^ 2) (↑(χ₈ 0) * τ ^ 0 + ↑(χ₈ 1) * τ ^ 1 + ↑(χ₈ 2) * τ ^ 2 + + ↑(χ₈ 3) * τ ^ 3 + ↑(χ₈ 4) * τ ^ 4 + ↑(χ₈ 5) * τ ^ 5 + ↑(χ₈ 6) * τ ^ 6 + ↑(χ₈ 7) * τ ^ 7) = ↑8, + { simp only [←h₄, χ₈_apply, matrix.cons_val_zero, algebra_map.coe_zero, zero_mul, + matrix.cons_val_one, matrix.head_cons, algebra_map.coe_one, matrix.cons_vec_bit0_eq_alt0, + matrix.cons_vec_append, matrix.cons_vec_alt0, matrix.cons_vec_bit1_eq_alt1, + matrix.cons_vec_alt1, int.cast_neg] }, + simpa only [hχ, one_mul, card, gauss_sum, ← h₅, h₁] using h, }, -- this allows us to apply `card_pow_char_pow` to our situation have h := char.card_pow_char_pow hq ψ₈.char (ring_char FF) n hu hFF hg, diff --git a/src/ring_theory/polynomial/gauss_lemma.lean b/src/ring_theory/polynomial/gauss_lemma.lean index b4074084240df..ec630a2b49f90 100644 --- a/src/ring_theory/polynomial/gauss_lemma.lean +++ b/src/ring_theory/polynomial/gauss_lemma.lean @@ -51,7 +51,6 @@ theorem integral_closure.mem_lifts_of_monic_of_dvd_map {f : R[X]} (hf : f.monic) {g : K[X]} (hg : g.monic) (hd : g ∣ f.map (algebra_map R K)) : g ∈ lifts (algebra_map (integral_closure R K) K) := begin - haveI : is_scalar_tower R K g.splitting_field := splitting_field_aux.is_scalar_tower _ _ _, have := mem_lift_of_splits_of_roots_mem_range (integral_closure R g.splitting_field) ((splits_id_iff_splits _).2 $ splitting_field.splits g) (hg.map _) (λ a ha, (set_like.ext_iff.mp (integral_closure R g.splitting_field).range_algebra_map _).mpr $ @@ -67,7 +66,7 @@ begin refine multiset.mem_of_le (roots.le_of_dvd ((hf.map _).map _).ne_zero _) ha, { apply_instance }, { exact map_dvd (algebra_map K g.splitting_field) hd }, - { apply splitting_field_aux.is_scalar_tower }, + { apply_instance } end variables [is_domain R] [is_fraction_ring R K] From 4b05d3f4f0601dca8abf99c4ec99187682ed0bba Mon Sep 17 00:00:00 2001 From: Riccardo Brasca Date: Thu, 15 Jun 2023 10:09:42 +0000 Subject: [PATCH 1190/1291] chore(field_theory.finite.galois_field, number_theory.cyclotomic.basic): make splitting_field.algebra' an instance again (#19191) --- src/field_theory/finite/galois_field.lean | 2 +- src/number_theory/cyclotomic/basic.lean | 2 -- 2 files changed, 1 insertion(+), 3 deletions(-) diff --git a/src/field_theory/finite/galois_field.lean b/src/field_theory/finite/galois_field.lean index 2199496911643..8961eb821974a 100644 --- a/src/field_theory/finite/galois_field.lean +++ b/src/field_theory/finite/galois_field.lean @@ -152,7 +152,7 @@ begin rw [splits_iff_card_roots, h1, ←finset.card_def, finset.card_univ, h2, zmod.card], end -local attribute [-instance] splitting_field.algebra' +local attribute [-instance] zmod.algebra /-- A Galois field with exponent 1 is equivalent to `zmod` -/ def equiv_zmod_p : galois_field p 1 ≃ₐ[zmod p] (zmod p) := diff --git a/src/number_theory/cyclotomic/basic.lean b/src/number_theory/cyclotomic/basic.lean index b30c3a02e03b4..abb7208658c20 100644 --- a/src/number_theory/cyclotomic/basic.lean +++ b/src/number_theory/cyclotomic/basic.lean @@ -479,8 +479,6 @@ lemma splitting_field_X_pow_sub_one : is_splitting_field K L (X ^ (n : ℕ) - 1) exact λ _, X_pow_sub_C_ne_zero n.pos (1 : L) end } -local attribute [-instance] splitting_field.algebra' - /-- Any two `n`-th cyclotomic extensions are isomorphic. -/ def alg_equiv (L' : Type*) [field L'] [algebra K L'] [is_cyclotomic_extension {n} K L'] : L ≃ₐ[K] L' := From 88474d1b5af6d37c2ab728b757771bced7f5194c Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Thu, 15 Jun 2023 13:54:35 +0000 Subject: [PATCH 1191/1291] chore(algebraic_geometry/Scheme): remove @[simps] from Spec (#19188) The `@[simps]` on `Spec` produces lemmas that aren't needed in mathlib3, and get in the way in mathlib4. This backports the removal of this attribute, to match https://github.com/leanprover-community/mathlib4/pull/5040. Co-authored-by: Scott Morrison Co-authored-by: Eric Wieser --- src/algebraic_geometry/AffineScheme.lean | 2 +- src/algebraic_geometry/Scheme.lean | 6 +++++- src/algebraic_geometry/properties.lean | 11 ++++++----- 3 files changed, 12 insertions(+), 7 deletions(-) diff --git a/src/algebraic_geometry/AffineScheme.lean b/src/algebraic_geometry/AffineScheme.lean index 59201acbd1616..4def63a940a0a 100644 --- a/src/algebraic_geometry/AffineScheme.lean +++ b/src/algebraic_geometry/AffineScheme.lean @@ -257,7 +257,7 @@ end lemma Scheme.Spec_map_presheaf_map_eq_to_hom {X : Scheme} {U V : opens X.carrier} (h : U = V) (W) : (Scheme.Spec.map (X.presheaf.map (eq_to_hom h).op).op).val.c.app W = - eq_to_hom (by { cases h, induction W using opposite.rec, dsimp, simp, refl }) := + eq_to_hom (by { cases h, induction W using opposite.rec, dsimp, simp, }) := begin have : Scheme.Spec.map (X.presheaf.map (𝟙 (op U))).op = 𝟙 _, { rw [X.presheaf.map_id, op_id, Scheme.Spec.map_id] }, diff --git a/src/algebraic_geometry/Scheme.lean b/src/algebraic_geometry/Scheme.lean index 888440d98ab1e..04f5850f5b0f0 100644 --- a/src/algebraic_geometry/Scheme.lean +++ b/src/algebraic_geometry/Scheme.lean @@ -158,7 +158,11 @@ Spec.LocallyRingedSpace_map_comp f g /-- The spectrum, as a contravariant functor from commutative rings to schemes. -/ -@[simps] def Spec : CommRingᵒᵖ ⥤ Scheme := +-- TODO: make either `Spec_obj` or `Spec.obj` the simp-normal form. `LocallyRingedSpace_obj` is +-- the simp-normal form of `toLocallyRingedSpace.obj`, but adding `simps` here without `attrs := []` +-- for the same effect caused problems in mathlib4. +@[simps {attrs := []}] +def Spec : CommRingᵒᵖ ⥤ Scheme := { obj := λ R, Spec_obj (unop R), map := λ R S f, Spec_map f.unop, map_id' := λ R, by rw [unop_id, Spec_map_id], diff --git a/src/algebraic_geometry/properties.lean b/src/algebraic_geometry/properties.lean index a67a35022dfcf..62429e066242e 100644 --- a/src/algebraic_geometry/properties.lean +++ b/src/algebraic_geometry/properties.lean @@ -308,14 +308,15 @@ begin Y.presheaf.obj _ ≅ _).symm.CommRing_iso_to_ring_equiv.is_domain _ end -instance {R : CommRing} [H : is_domain R] : is_integral (Scheme.Spec.obj $ op R) := +instance {R : CommRing} [H : is_domain R] : irreducible_space (Scheme.Spec.obj $ op R).carrier := begin - apply_with is_integral_of_is_irreducible_is_reduced { instances := ff }, - { apply_instance }, - { dsimp [Spec.Top_obj], - apply_instance }, + convert prime_spectrum.irreducible_space, + assumption end +instance {R : CommRing} [is_domain R] : is_integral (Scheme.Spec.obj $ op R) := +is_integral_of_is_irreducible_is_reduced _ + lemma affine_is_integral_iff (R : CommRing) : is_integral (Scheme.Spec.obj $ op R) ↔ is_domain R := ⟨λ h, by exactI ring_equiv.is_domain ((Scheme.Spec.obj $ op R).presheaf.obj _) From bf9bbbcf0c1c1ead18280b0d010e417b10abb1b6 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 15 Jun 2023 15:45:26 +0000 Subject: [PATCH 1192/1291] feat(field_theory/ratfunc): The numerator and denominator of a rational function are coprime (#18652) Also make more arguments to `gcd_ne_zero_of_left`/`gcd_ne_zero_of_right` implicit. Co-authored-by: Bhavik Mehta --- src/algebra/euclidean_domain/basic.lean | 3 +++ src/field_theory/ratfunc.lean | 8 ++++++++ src/ring_theory/euclidean_domain.lean | 15 ++++++++++----- 3 files changed, 21 insertions(+), 5 deletions(-) diff --git a/src/algebra/euclidean_domain/basic.lean b/src/algebra/euclidean_domain/basic.lean index fde4f36389841..026e3911720ba 100644 --- a/src/algebra/euclidean_domain/basic.lean +++ b/src/algebra/euclidean_domain/basic.lean @@ -84,6 +84,9 @@ begin rw [mul_div_cancel_left _ hz, mul_left_comm, mul_div_cancel_left _ hz] end +protected lemma mul_div_cancel' {a b : R} (hb : b ≠ 0) (hab : b ∣ a) : b * (a / b) = a := +by rw [←mul_div_assoc _ hab, mul_div_cancel_left _ hb] + @[simp, priority 900] -- This generalizes `int.div_one`, see note [simp-normal form] lemma div_one (p : R) : p / 1 = p := (euclidean_domain.eq_div_of_mul_eq_left (one_ne_zero' R) (mul_one p)).symm diff --git a/src/field_theory/ratfunc.lean b/src/field_theory/ratfunc.lean index e56b902d45683..735db2340429f 100644 --- a/src/field_theory/ratfunc.lean +++ b/src/field_theory/ratfunc.lean @@ -1030,6 +1030,14 @@ x.induction_on (λ p q hq, begin exact inv_ne_zero (polynomial.leading_coeff_ne_zero.mpr q_div_ne_zero) }, end) +lemma is_coprime_num_denom (x : ratfunc K) : is_coprime x.num x.denom := +begin + induction x using ratfunc.induction_on with p q hq, + rw [num_div, denom_div _ hq], + exact (is_coprime_mul_unit_left ((leading_coeff_ne_zero.2 $ right_div_gcd_ne_zero + hq).is_unit.inv.map C) _ _).2 (is_coprime_div_gcd_div_gcd hq), +end + @[simp] lemma num_eq_zero_iff {x : ratfunc K} : num x = 0 ↔ x = 0 := ⟨λ h, by rw [← num_div_denom x, h, ring_hom.map_zero, zero_div], λ h, h.symm ▸ num_zero⟩ diff --git a/src/ring_theory/euclidean_domain.lean b/src/ring_theory/euclidean_domain.lean index 7ed33ee3b4815..c8a290d8ffc03 100644 --- a/src/ring_theory/euclidean_domain.lean +++ b/src/ring_theory/euclidean_domain.lean @@ -29,14 +29,12 @@ open euclidean_domain set ideal section gcd_monoid -variables {R : Type*} [euclidean_domain R] [gcd_monoid R] +variables {R : Type*} [euclidean_domain R] [gcd_monoid R] {p q : R} -lemma gcd_ne_zero_of_left (p q : R) (hp : p ≠ 0) : - gcd_monoid.gcd p q ≠ 0 := +lemma gcd_ne_zero_of_left (hp : p ≠ 0) : gcd_monoid.gcd p q ≠ 0 := λ h, hp $ eq_zero_of_zero_dvd (h ▸ gcd_dvd_left p q) -lemma gcd_ne_zero_of_right (p q : R) (hp : q ≠ 0) : - gcd_monoid.gcd p q ≠ 0 := +lemma gcd_ne_zero_of_right (hp : q ≠ 0) : gcd_monoid.gcd p q ≠ 0 := λ h, hp $ eq_zero_of_zero_dvd (h ▸ gcd_dvd_right p q) lemma left_div_gcd_ne_zero {p q : R} (hp : p ≠ 0) : @@ -57,6 +55,13 @@ begin exact r0, end +lemma is_coprime_div_gcd_div_gcd (hq : q ≠ 0) : + is_coprime (p / gcd_monoid.gcd p q) (q / gcd_monoid.gcd p q) := +(gcd_is_unit_iff _ _).1 $ is_unit_gcd_of_eq_mul_gcd + (euclidean_domain.mul_div_cancel' (gcd_ne_zero_of_right hq) $ gcd_dvd_left _ _).symm + (euclidean_domain.mul_div_cancel' (gcd_ne_zero_of_right hq) $ gcd_dvd_right _ _).symm $ + gcd_ne_zero_of_right hq + end gcd_monoid namespace euclidean_domain From 5b2fe80501ff327b9109fb09b7cc8c325cd0d7d9 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Fri, 16 Jun 2023 04:21:23 +0000 Subject: [PATCH 1193/1291] chore(number_theory/legendre_symbol/quadratic_reciprocity): split some files (#19193) The best split here is separating `quadratic_reciprocity` into that and `legendre_symbol`. I was looking at this because the olean analyser in the port-progress-bot made an unlikely claim about unnecessary dependencies, and I wanted to make sure we weren't missing something in doubting it. We weren't. :-) Co-authored-by: Scott Morrison --- src/number_theory/legendre_symbol/basic.lean | 298 ++++++++++++++++++ .../basic.lean} | 116 ------- .../quadratic_char/gauss_sum.lean | 145 +++++++++ .../quadratic_reciprocity.lean | 270 +--------------- src/number_theory/sum_two_squares.lean | 2 +- src/number_theory/zsqrtd/gaussian_int.lean | 78 +---- .../zsqrtd/quadratic_reciprocity.lean | 96 ++++++ 7 files changed, 550 insertions(+), 455 deletions(-) create mode 100644 src/number_theory/legendre_symbol/basic.lean rename src/number_theory/legendre_symbol/{quadratic_char.lean => quadratic_char/basic.lean} (69%) create mode 100644 src/number_theory/legendre_symbol/quadratic_char/gauss_sum.lean create mode 100644 src/number_theory/zsqrtd/quadratic_reciprocity.lean diff --git a/src/number_theory/legendre_symbol/basic.lean b/src/number_theory/legendre_symbol/basic.lean new file mode 100644 index 0000000000000..545e7e6149a45 --- /dev/null +++ b/src/number_theory/legendre_symbol/basic.lean @@ -0,0 +1,298 @@ +/- +Copyright (c) 2018 Chris Hughes. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Chris Hughes, Michael Stoll +-/ +import number_theory.legendre_symbol.quadratic_char.basic + +/-! +# Legendre symbol + +This file contains results about Legendre symbols. + +We define the Legendre symbol $\Bigl(\frac{a}{p}\Bigr)$ as `legendre_sym p a`. +Note the order of arguments! The advantage of this form is that then `legendre_sym p` +is a multiplicative map. + +The Legendre symbol is used to define the Jacobi symbol, `jacobi_sym a b`, for integers `a` +and (odd) natural numbers `b`, which extends the Legendre symbol. + +## Main results + +We also prove the supplementary laws that give conditions for when `-1` +is a square modulo a prime `p`: +`legendre_sym.at_neg_one` and `zmod.exists_sq_eq_neg_one_iff` for `-1`. + +See `number_theory.legendre_symbol.quadratic_reciprocity` for the conditions when `2` and `-2` +are squares: +`legendre_sym.at_two` and `zmod.exists_sq_eq_two_iff` for `2`, +`legendre_sym.at_neg_two` and `zmod.exists_sq_eq_neg_two_iff` for `-2`. + +## Tags + +quadratic residue, quadratic nonresidue, Legendre symbol +-/ + +open nat + +section euler + +namespace zmod + +variables (p : ℕ) [fact p.prime] + +/-- Euler's Criterion: A unit `x` of `zmod p` is a square if and only if `x ^ (p / 2) = 1`. -/ +lemma euler_criterion_units (x : (zmod p)ˣ) : (∃ y : (zmod p)ˣ, y ^ 2 = x) ↔ x ^ (p / 2) = 1 := +begin + by_cases hc : p = 2, + { substI hc, + simp only [eq_iff_true_of_subsingleton, exists_const], }, + { have h₀ := finite_field.unit_is_square_iff (by rwa ring_char_zmod_n) x, + have hs : (∃ y : (zmod p)ˣ, y ^ 2 = x) ↔ is_square(x) := + by { rw is_square_iff_exists_sq x, + simp_rw eq_comm, }, + rw hs, + rwa card p at h₀, }, +end + +/-- Euler's Criterion: a nonzero `a : zmod p` is a square if and only if `x ^ (p / 2) = 1`. -/ +lemma euler_criterion {a : zmod p} (ha : a ≠ 0) : is_square (a : zmod p) ↔ a ^ (p / 2) = 1 := +begin + apply (iff_congr _ (by simp [units.ext_iff])).mp (euler_criterion_units p (units.mk0 a ha)), + simp only [units.ext_iff, sq, units.coe_mk0, units.coe_mul], + split, { rintro ⟨y, hy⟩, exact ⟨y, hy.symm⟩ }, + { rintro ⟨y, rfl⟩, + have hy : y ≠ 0, { rintro rfl, simpa [zero_pow] using ha, }, + refine ⟨units.mk0 y hy, _⟩, simp, } +end + +/-- If `a : zmod p` is nonzero, then `a^(p/2)` is either `1` or `-1`. -/ +lemma pow_div_two_eq_neg_one_or_one {a : zmod p} (ha : a ≠ 0) : + a ^ (p / 2) = 1 ∨ a ^ (p / 2) = -1 := +begin + cases prime.eq_two_or_odd (fact.out p.prime) with hp2 hp_odd, + { substI p, revert a ha, dec_trivial }, + rw [← mul_self_eq_one_iff, ← pow_add, ← two_mul, two_mul_odd_div_two hp_odd], + exact pow_card_sub_one_eq_one ha +end + +end zmod + +end euler + +section legendre + +/-! +### Definition of the Legendre symbol and basic properties +-/ + +open zmod + +variables (p : ℕ) [fact p.prime] + +/-- The Legendre symbol of `a : ℤ` and a prime `p`, `legendre_sym p a`, +is an integer defined as + +* `0` if `a` is `0` modulo `p`; +* `1` if `a` is a nonzero square modulo `p` +* `-1` otherwise. + +Note the order of the arguments! The advantage of the order chosen here is +that `legendre_sym p` is a multiplicative function `ℤ → ℤ`. +-/ +def legendre_sym (a : ℤ) : ℤ := quadratic_char (zmod p) a + +namespace legendre_sym + +/-- We have the congruence `legendre_sym p a ≡ a ^ (p / 2) mod p`. -/ +lemma eq_pow (a : ℤ) : (legendre_sym p a : zmod p) = a ^ (p / 2) := +begin + cases eq_or_ne (ring_char (zmod p)) 2 with hc hc, + { by_cases ha : (a : zmod p) = 0, + { rw [legendre_sym, ha, quadratic_char_zero, + zero_pow (nat.div_pos (fact.out p.prime).two_le (succ_pos 1))], + norm_cast, }, + { have := (ring_char_zmod_n p).symm.trans hc, -- p = 2 + substI p, + rw [legendre_sym, quadratic_char_eq_one_of_char_two hc ha], + revert ha, + generalize : (a : zmod 2) = b, revert b, dec_trivial } }, + { convert quadratic_char_eq_pow_of_char_ne_two' hc (a : zmod p), + exact (card p).symm }, +end + +/-- If `p ∤ a`, then `legendre_sym p a` is `1` or `-1`. -/ +lemma eq_one_or_neg_one {a : ℤ} (ha : (a : zmod p) ≠ 0) : + legendre_sym p a = 1 ∨ legendre_sym p a = -1 := +quadratic_char_dichotomy ha + +lemma 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 eq_zero_iff (a : ℤ) : legendre_sym p a = 0 ↔ (a : zmod p) = 0 := +quadratic_char_eq_zero_iff + +@[simp] lemma at_zero : legendre_sym p 0 = 0 := +by rw [legendre_sym, int.cast_zero, mul_char.map_zero] + +@[simp] lemma at_one : legendre_sym p 1 = 1 := +by rw [legendre_sym, int.cast_one, mul_char.map_one] + +/-- The Legendre symbol is multiplicative in `a` for `p` fixed. -/ +protected +lemma mul (a b : ℤ) : legendre_sym p (a * b) = legendre_sym p a * legendre_sym p b := +by simp only [legendre_sym, int.cast_mul, map_mul] + +/-- The Legendre symbol is a homomorphism of monoids with zero. -/ +@[simps] def hom : ℤ →*₀ ℤ := +{ to_fun := legendre_sym p, + map_zero' := at_zero p, + map_one' := at_one p, + map_mul' := legendre_sym.mul p } + +/-- The square of the symbol is 1 if `p ∤ a`. -/ +theorem sq_one {a : ℤ} (ha : (a : zmod p) ≠ 0) : (legendre_sym p a) ^ 2 = 1 := +quadratic_char_sq_one ha + +/-- The Legendre symbol of `a^2` at `p` is 1 if `p ∤ a`. -/ +theorem sq_one' {a : ℤ} (ha : (a : zmod p) ≠ 0) : legendre_sym p (a ^ 2) = 1 := +by exact_mod_cast quadratic_char_sq_one' ha + +/-- The Legendre symbol depends only on `a` mod `p`. -/ +protected +theorem mod (a : ℤ) : legendre_sym p a = legendre_sym p (a % p) := +by simp only [legendre_sym, int_cast_mod] + +/-- When `p ∤ a`, then `legendre_sym p a = 1` iff `a` is a square mod `p`. -/ +lemma eq_one_iff {a : ℤ} (ha0 : (a : zmod p) ≠ 0) : + legendre_sym p a = 1 ↔ is_square (a : zmod p) := +quadratic_char_one_iff_is_square ha0 + +lemma eq_one_iff' {a : ℕ} (ha0 : (a : zmod p) ≠ 0) : + legendre_sym p a = 1 ↔ is_square (a : zmod p) := +by {rw eq_one_iff, norm_cast, exact_mod_cast ha0} + +/-- `legendre_sym p a = -1` iff `a` is a nonsquare mod `p`. -/ +lemma eq_neg_one_iff {a : ℤ} : legendre_sym p a = -1 ↔ ¬ is_square (a : zmod p) := +quadratic_char_neg_one_iff_not_is_square + +lemma eq_neg_one_iff' {a : ℕ} : legendre_sym p a = -1 ↔ ¬ is_square (a : zmod p) := +by {rw eq_neg_one_iff, norm_cast} + +/-- The number of square roots of `a` modulo `p` is determined by the Legendre symbol. -/ +lemma card_sqrts (hp : p ≠ 2) (a : ℤ) : + ↑{x : zmod p | x^2 = a}.to_finset.card = legendre_sym p a + 1 := +quadratic_char_card_sqrts ((ring_char_zmod_n p).substr hp) a + +end legendre_sym + +end legendre + +section quadratic_form + +/-! +### Applications to binary quadratic forms +-/ + +namespace legendre_sym + +/-- The Legendre symbol `legendre_sym p a = 1` if there is a solution in `ℤ/pℤ` +of the equation `x^2 - a*y^2 = 0` with `y ≠ 0`. -/ +lemma eq_one_of_sq_sub_mul_sq_eq_zero {p : ℕ} [fact p.prime] + {a : ℤ} (ha : (a : zmod p) ≠ 0) {x y : zmod p} (hy : y ≠ 0) (hxy : x ^ 2 - a * y ^ 2 = 0) : + legendre_sym p a = 1 := +begin + apply_fun (* y⁻¹ ^ 2) at hxy, + simp only [zero_mul] at hxy, + rw [(by ring : (x ^ 2 - ↑a * y ^ 2) * y⁻¹ ^ 2 = (x * y⁻¹) ^ 2 - a * (y * y⁻¹) ^ 2), + mul_inv_cancel hy, one_pow, mul_one, sub_eq_zero, pow_two] at hxy, + exact (eq_one_iff p ha).mpr ⟨x * y⁻¹, hxy.symm⟩, +end + +/-- The Legendre symbol `legendre_sym p a = 1` if there is a solution in `ℤ/pℤ` +of the equation `x^2 - a*y^2 = 0` with `x ≠ 0`. -/ +lemma eq_one_of_sq_sub_mul_sq_eq_zero' {p : ℕ} [fact p.prime] + {a : ℤ} (ha : (a : zmod p) ≠ 0) {x y : zmod p} (hx : x ≠ 0) (hxy : x ^ 2 - a * y ^ 2 = 0) : + legendre_sym p a = 1 := +begin + have hy : y ≠ 0, + { rintro rfl, + rw [zero_pow' 2 (by norm_num), mul_zero, sub_zero, pow_eq_zero_iff (by norm_num : 0 < 2)] + at hxy, + exacts [hx hxy, infer_instance], }, -- why is the instance not inferred automatically? + exact eq_one_of_sq_sub_mul_sq_eq_zero ha hy hxy, +end + +/-- If `legendre_sym p a = -1`, then the only solution of `x^2 - a*y^2 = 0` in `ℤ/pℤ` +is the trivial one. -/ +lemma eq_zero_mod_of_eq_neg_one {p : ℕ} [fact p.prime] {a : ℤ} + (h : legendre_sym p a = -1) {x y : zmod p} (hxy : x ^ 2 - a * y ^ 2 = 0) : x = 0 ∧ y = 0 := +begin + have ha : (a : zmod p) ≠ 0, + { intro hf, + rw (eq_zero_iff p a).mpr hf at h, + exact int.zero_ne_neg_of_ne zero_ne_one h, }, + by_contra hf, + cases not_and_distrib.mp hf with hx hy, + { rw [eq_one_of_sq_sub_mul_sq_eq_zero' ha hx hxy, eq_neg_self_iff] at h, + exact one_ne_zero h, }, + { rw [eq_one_of_sq_sub_mul_sq_eq_zero ha hy hxy, eq_neg_self_iff] at h, + exact one_ne_zero h, } +end + +/-- If `legendre_sym p a = -1` and `p` divides `x^2 - a*y^2`, then `p` must divide `x` and `y`. -/ +lemma prime_dvd_of_eq_neg_one {p : ℕ} [fact p.prime] {a : ℤ} + (h : legendre_sym p a = -1) {x y : ℤ} (hxy : ↑p ∣ x ^ 2 - a * y ^ 2) : ↑p ∣ x ∧ ↑p ∣ y := +begin + simp_rw ← zmod.int_coe_zmod_eq_zero_iff_dvd at hxy ⊢, + push_cast at hxy, + exact eq_zero_mod_of_eq_neg_one h hxy, +end + +end legendre_sym + +end quadratic_form + +section values + +/-! +### The value of the Legendre symbol at `-1` + +See `jacobi_sym.at_neg_one` for the corresponding statement for the Jacobi symbol. +-/ + +variables {p : ℕ} [fact p.prime] + +open zmod + +/-- `legendre_sym p (-1)` is given by `χ₄ p`. -/ +lemma legendre_sym.at_neg_one (hp : p ≠ 2) : legendre_sym p (-1) = χ₄ p := +by simp only [legendre_sym, card p, quadratic_char_neg_one ((ring_char_zmod_n p).substr hp), + int.cast_neg, int.cast_one] + +namespace zmod + +/-- `-1` is a square in `zmod p` iff `p` is not congruent to `3` mod `4`. -/ +lemma exists_sq_eq_neg_one_iff : is_square (-1 : zmod p) ↔ p % 4 ≠ 3 := +by rw [finite_field.is_square_neg_one_iff, card p] + +lemma mod_four_ne_three_of_sq_eq_neg_one {y : zmod p} (hy : y ^ 2 = -1) : p % 4 ≠ 3 := +exists_sq_eq_neg_one_iff.1 ⟨y, hy ▸ pow_two y⟩ + +/-- If two nonzero squares are negatives of each other in `zmod p`, then `p % 4 ≠ 3`. -/ +lemma mod_four_ne_three_of_sq_eq_neg_sq' {x y : zmod p} (hy : y ≠ 0) (hxy : x ^ 2 = - y ^ 2) : + p % 4 ≠ 3 := +@mod_four_ne_three_of_sq_eq_neg_one p _ (x / y) begin + apply_fun (λ z, z / y ^ 2) at hxy, + rwa [neg_div, ←div_pow, ←div_pow, div_self hy, one_pow] at hxy +end + +lemma mod_four_ne_three_of_sq_eq_neg_sq {x y : zmod p} (hx : x ≠ 0) (hxy : x ^ 2 = - y ^ 2) : + p % 4 ≠ 3 := +mod_four_ne_three_of_sq_eq_neg_sq' hx (neg_eq_iff_eq_neg.mpr hxy).symm + +end zmod + +end values diff --git a/src/number_theory/legendre_symbol/quadratic_char.lean b/src/number_theory/legendre_symbol/quadratic_char/basic.lean similarity index 69% rename from src/number_theory/legendre_symbol/quadratic_char.lean rename to src/number_theory/legendre_symbol/quadratic_char/basic.lean index b49542517b746..0d87db3f03a63 100644 --- a/src/number_theory/legendre_symbol/quadratic_char.lean +++ b/src/number_theory/legendre_symbol/quadratic_char/basic.lean @@ -6,7 +6,6 @@ Authors: Michael Stoll import data.fintype.parity import number_theory.legendre_symbol.zmod_char import field_theory.finite.basic -import number_theory.legendre_symbol.gauss_sum /-! # Quadratic characters of finite fields @@ -319,119 +318,4 @@ begin or.resolve_right (nat.odd_mod_four_iff.mp h₁)⟩, }, end -/-- The value of the quadratic character at `2` -/ -lemma quadratic_char_two [decidable_eq F] (hF : ring_char F ≠ 2) : - quadratic_char F 2 = χ₈ (fintype.card F) := -is_quadratic.eq_of_eq_coe (quadratic_char_is_quadratic F) is_quadratic_χ₈ hF - ((quadratic_char_eq_pow_of_char_ne_two' hF 2).trans (finite_field.two_pow_card hF)) - -/-- `2` is a square in `F` iff `#F` is not congruent to `3` or `5` mod `8`. -/ -lemma finite_field.is_square_two_iff : - is_square (2 : F) ↔ fintype.card F % 8 ≠ 3 ∧ fintype.card F % 8 ≠ 5 := -begin - classical, - by_cases hF : ring_char F = 2, - focus - { have h := finite_field.even_card_of_char_two hF, - simp only [finite_field.is_square_of_char_two hF, true_iff], }, - rotate, focus - { have h := finite_field.odd_card_of_char_ne_two hF, - rw [← quadratic_char_one_iff_is_square (ring.two_ne_zero hF), quadratic_char_two hF, - χ₈_nat_eq_if_mod_eight], - simp only [h, nat.one_ne_zero, if_false, ite_eq_left_iff, ne.def, (dec_trivial : (-1 : ℤ) ≠ 1), - imp_false, not_not], }, - all_goals - { rw [← nat.mod_mod_of_dvd _ (by norm_num : 2 ∣ 8)] at h, - have h₁ := nat.mod_lt (fintype.card F) (dec_trivial : 0 < 8), - revert h₁ h, - generalize : fintype.card F % 8 = n, - dec_trivial!, } -end - -/-- The value of the quadratic character at `-2` -/ -lemma quadratic_char_neg_two [decidable_eq F] (hF : ring_char F ≠ 2) : - quadratic_char F (-2) = χ₈' (fintype.card F) := -begin - rw [(by norm_num : (-2 : F) = (-1) * 2), map_mul, χ₈'_eq_χ₄_mul_χ₈, quadratic_char_neg_one hF, - quadratic_char_two hF, @cast_nat_cast _ (zmod 4) _ _ _ (by norm_num : 4 ∣ 8)], -end - -/-- `-2` is a square in `F` iff `#F` is not congruent to `5` or `7` mod `8`. -/ -lemma finite_field.is_square_neg_two_iff : - is_square (-2 : F) ↔ fintype.card F % 8 ≠ 5 ∧ fintype.card F % 8 ≠ 7 := -begin - classical, - by_cases hF : ring_char F = 2, - focus - { have h := finite_field.even_card_of_char_two hF, - simp only [finite_field.is_square_of_char_two hF, true_iff], }, - rotate, focus - { have h := finite_field.odd_card_of_char_ne_two hF, - rw [← quadratic_char_one_iff_is_square (neg_ne_zero.mpr (ring.two_ne_zero hF)), - quadratic_char_neg_two hF, χ₈'_nat_eq_if_mod_eight], - simp only [h, nat.one_ne_zero, if_false, ite_eq_left_iff, ne.def, (dec_trivial : (-1 : ℤ) ≠ 1), - imp_false, not_not], }, - all_goals - { rw [← nat.mod_mod_of_dvd _ (by norm_num : 2 ∣ 8)] at h, - have h₁ := nat.mod_lt (fintype.card F) (dec_trivial : 0 < 8), - revert h₁ h, - generalize : fintype.card F % 8 = n, - dec_trivial! } -end - -/-- The relation between the values of the quadratic character of one field `F` at the -cardinality of another field `F'` and of the quadratic character of `F'` at the cardinality -of `F`. -/ -lemma quadratic_char_card_card [decidable_eq F] (hF : ring_char F ≠ 2) {F' : Type*} [field F'] - [fintype F'] [decidable_eq F'] (hF' : ring_char F' ≠ 2) (h : ring_char F' ≠ ring_char F) : - quadratic_char F (fintype.card F') = quadratic_char F' (quadratic_char F (-1) * fintype.card F) := -begin - let χ := (quadratic_char F).ring_hom_comp (algebra_map ℤ F'), - have hχ₁ : χ.is_nontrivial, - { obtain ⟨a, ha⟩ := quadratic_char_exists_neg_one hF, - have hu : is_unit a, - { contrapose ha, - exact ne_of_eq_of_ne (map_nonunit (quadratic_char F) ha) - (mt zero_eq_neg.mp one_ne_zero), }, - use hu.unit, - simp only [is_unit.unit_spec, ring_hom_comp_apply, eq_int_cast, ne.def, ha], - rw [int.cast_neg, int.cast_one], - exact ring.neg_one_ne_one_of_char_ne_two hF', }, - have hχ₂ : χ.is_quadratic := is_quadratic.comp (quadratic_char_is_quadratic F) _, - have h := char.card_pow_card hχ₁ hχ₂ h hF', - rw [← quadratic_char_eq_pow_of_char_ne_two' hF'] at h, - exact (is_quadratic.eq_of_eq_coe (quadratic_char_is_quadratic F') - (quadratic_char_is_quadratic F) hF' h).symm, -end - -/-- The value of the quadratic character at an odd prime `p` different from `ring_char F`. -/ -lemma quadratic_char_odd_prime [decidable_eq F] (hF : ring_char F ≠ 2) {p : ℕ} [fact p.prime] - (hp₁ : p ≠ 2) (hp₂ : ring_char F ≠ p) : - quadratic_char F p = quadratic_char (zmod p) (χ₄ (fintype.card F) * fintype.card F) := -begin - rw [← quadratic_char_neg_one hF], - have h := quadratic_char_card_card hF (ne_of_eq_of_ne (ring_char_zmod_n p) hp₁) - (ne_of_eq_of_ne (ring_char_zmod_n p) hp₂.symm), - rwa [card p] at h, -end - -/-- An odd prime `p` is a square in `F` iff the quadratic character of `zmod p` does not -take the value `-1` on `χ₄(#F) * #F`. -/ -lemma finite_field.is_square_odd_prime_iff (hF : ring_char F ≠ 2) {p : ℕ} [fact p.prime] - (hp : p ≠ 2) : - is_square (p : F) ↔ quadratic_char (zmod p) (χ₄ (fintype.card F) * fintype.card F) ≠ -1 := -begin - classical, - by_cases hFp : ring_char F = p, - { rw [show (p : F) = 0, by { rw ← hFp, exact ring_char.nat.cast_ring_char }], - simp only [is_square_zero, ne.def, true_iff, map_mul], - obtain ⟨n, _, hc⟩ := finite_field.card F (ring_char F), - have hchar : ring_char F = ring_char (zmod p) := by {rw hFp, exact (ring_char_zmod_n p).symm}, - conv {congr, to_lhs, congr, skip, rw [hc, nat.cast_pow, map_pow, hchar, map_ring_char], }, - simp only [zero_pow n.pos, mul_zero, zero_eq_neg, one_ne_zero, not_false_iff], }, - { rw [← iff.not_left (@quadratic_char_neg_one_iff_not_is_square F _ _ _ _), - quadratic_char_odd_prime hF hp], - exact hFp, }, -end - end special_values diff --git a/src/number_theory/legendre_symbol/quadratic_char/gauss_sum.lean b/src/number_theory/legendre_symbol/quadratic_char/gauss_sum.lean new file mode 100644 index 0000000000000..6e220b86a5d81 --- /dev/null +++ b/src/number_theory/legendre_symbol/quadratic_char/gauss_sum.lean @@ -0,0 +1,145 @@ +/- +Copyright (c) 2022 Michael Stoll. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Michael Stoll +-/ +import number_theory.legendre_symbol.quadratic_char.basic +import number_theory.legendre_symbol.gauss_sum + +/-! +# Quadratic characters of finite fields + +Further facts relying on Gauss sums. + +-/ + +/-! +### Basic properties of the quadratic character + +We prove some properties of the quadratic character. +We work with a finite field `F` here. +The interesting case is when the characteristic of `F` is odd. +-/ + +section special_values + +open zmod mul_char + +variables {F : Type*} [field F] [fintype F] + +/-- The value of the quadratic character at `2` -/ +lemma quadratic_char_two [decidable_eq F] (hF : ring_char F ≠ 2) : + quadratic_char F 2 = χ₈ (fintype.card F) := +is_quadratic.eq_of_eq_coe (quadratic_char_is_quadratic F) is_quadratic_χ₈ hF + ((quadratic_char_eq_pow_of_char_ne_two' hF 2).trans (finite_field.two_pow_card hF)) + +/-- `2` is a square in `F` iff `#F` is not congruent to `3` or `5` mod `8`. -/ +lemma finite_field.is_square_two_iff : + is_square (2 : F) ↔ fintype.card F % 8 ≠ 3 ∧ fintype.card F % 8 ≠ 5 := +begin + classical, + by_cases hF : ring_char F = 2, + focus + { have h := finite_field.even_card_of_char_two hF, + simp only [finite_field.is_square_of_char_two hF, true_iff], }, + rotate, focus + { have h := finite_field.odd_card_of_char_ne_two hF, + rw [← quadratic_char_one_iff_is_square (ring.two_ne_zero hF), quadratic_char_two hF, + χ₈_nat_eq_if_mod_eight], + simp only [h, nat.one_ne_zero, if_false, ite_eq_left_iff, ne.def, (dec_trivial : (-1 : ℤ) ≠ 1), + imp_false, not_not], }, + all_goals + { rw [← nat.mod_mod_of_dvd _ (by norm_num : 2 ∣ 8)] at h, + have h₁ := nat.mod_lt (fintype.card F) (dec_trivial : 0 < 8), + revert h₁ h, + generalize : fintype.card F % 8 = n, + dec_trivial!, } +end + +/-- The value of the quadratic character at `-2` -/ +lemma quadratic_char_neg_two [decidable_eq F] (hF : ring_char F ≠ 2) : + quadratic_char F (-2) = χ₈' (fintype.card F) := +begin + rw [(by norm_num : (-2 : F) = (-1) * 2), map_mul, χ₈'_eq_χ₄_mul_χ₈, quadratic_char_neg_one hF, + quadratic_char_two hF, @cast_nat_cast _ (zmod 4) _ _ _ (by norm_num : 4 ∣ 8)], +end + +/-- `-2` is a square in `F` iff `#F` is not congruent to `5` or `7` mod `8`. -/ +lemma finite_field.is_square_neg_two_iff : + is_square (-2 : F) ↔ fintype.card F % 8 ≠ 5 ∧ fintype.card F % 8 ≠ 7 := +begin + classical, + by_cases hF : ring_char F = 2, + focus + { have h := finite_field.even_card_of_char_two hF, + simp only [finite_field.is_square_of_char_two hF, true_iff], }, + rotate, focus + { have h := finite_field.odd_card_of_char_ne_two hF, + rw [← quadratic_char_one_iff_is_square (neg_ne_zero.mpr (ring.two_ne_zero hF)), + quadratic_char_neg_two hF, χ₈'_nat_eq_if_mod_eight], + simp only [h, nat.one_ne_zero, if_false, ite_eq_left_iff, ne.def, (dec_trivial : (-1 : ℤ) ≠ 1), + imp_false, not_not], }, + all_goals + { rw [← nat.mod_mod_of_dvd _ (by norm_num : 2 ∣ 8)] at h, + have h₁ := nat.mod_lt (fintype.card F) (dec_trivial : 0 < 8), + revert h₁ h, + generalize : fintype.card F % 8 = n, + dec_trivial! } +end + +/-- The relation between the values of the quadratic character of one field `F` at the +cardinality of another field `F'` and of the quadratic character of `F'` at the cardinality +of `F`. -/ +lemma quadratic_char_card_card [decidable_eq F] (hF : ring_char F ≠ 2) {F' : Type*} [field F'] + [fintype F'] [decidable_eq F'] (hF' : ring_char F' ≠ 2) (h : ring_char F' ≠ ring_char F) : + quadratic_char F (fintype.card F') = quadratic_char F' (quadratic_char F (-1) * fintype.card F) := +begin + let χ := (quadratic_char F).ring_hom_comp (algebra_map ℤ F'), + have hχ₁ : χ.is_nontrivial, + { obtain ⟨a, ha⟩ := quadratic_char_exists_neg_one hF, + have hu : is_unit a, + { contrapose ha, + exact ne_of_eq_of_ne (map_nonunit (quadratic_char F) ha) + (mt zero_eq_neg.mp one_ne_zero), }, + use hu.unit, + simp only [is_unit.unit_spec, ring_hom_comp_apply, eq_int_cast, ne.def, ha], + rw [int.cast_neg, int.cast_one], + exact ring.neg_one_ne_one_of_char_ne_two hF', }, + have hχ₂ : χ.is_quadratic := is_quadratic.comp (quadratic_char_is_quadratic F) _, + have h := char.card_pow_card hχ₁ hχ₂ h hF', + rw [← quadratic_char_eq_pow_of_char_ne_two' hF'] at h, + exact (is_quadratic.eq_of_eq_coe (quadratic_char_is_quadratic F') + (quadratic_char_is_quadratic F) hF' h).symm, +end + +/-- The value of the quadratic character at an odd prime `p` different from `ring_char F`. -/ +lemma quadratic_char_odd_prime [decidable_eq F] (hF : ring_char F ≠ 2) {p : ℕ} [fact p.prime] + (hp₁ : p ≠ 2) (hp₂ : ring_char F ≠ p) : + quadratic_char F p = quadratic_char (zmod p) (χ₄ (fintype.card F) * fintype.card F) := +begin + rw [← quadratic_char_neg_one hF], + have h := quadratic_char_card_card hF (ne_of_eq_of_ne (ring_char_zmod_n p) hp₁) + (ne_of_eq_of_ne (ring_char_zmod_n p) hp₂.symm), + rwa [card p] at h, +end + +/-- An odd prime `p` is a square in `F` iff the quadratic character of `zmod p` does not +take the value `-1` on `χ₄(#F) * #F`. -/ +lemma finite_field.is_square_odd_prime_iff (hF : ring_char F ≠ 2) {p : ℕ} [fact p.prime] + (hp : p ≠ 2) : + is_square (p : F) ↔ quadratic_char (zmod p) (χ₄ (fintype.card F) * fintype.card F) ≠ -1 := +begin + classical, + by_cases hFp : ring_char F = p, + { rw [show (p : F) = 0, by { rw ← hFp, exact ring_char.nat.cast_ring_char }], + simp only [is_square_zero, ne.def, true_iff, map_mul], + obtain ⟨n, _, hc⟩ := finite_field.card F (ring_char F), + have hchar : ring_char F = ring_char (zmod p) := by {rw hFp, exact (ring_char_zmod_n p).symm}, + conv {congr, to_lhs, congr, skip, rw [hc, nat.cast_pow, map_pow, hchar, map_ring_char], }, + simp only [zero_pow n.pos, mul_zero, zero_eq_neg, one_ne_zero, not_false_iff], }, + { rw [← iff.not_left (@quadratic_char_neg_one_iff_not_is_square F _ _ _ _), + quadratic_char_odd_prime hF hp], + exact hFp, }, +end + +end special_values diff --git a/src/number_theory/legendre_symbol/quadratic_reciprocity.lean b/src/number_theory/legendre_symbol/quadratic_reciprocity.lean index 946ff967ed92d..53cfb747b8efe 100644 --- a/src/number_theory/legendre_symbol/quadratic_reciprocity.lean +++ b/src/number_theory/legendre_symbol/quadratic_reciprocity.lean @@ -3,19 +3,11 @@ Copyright (c) 2018 Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes, Michael Stoll -/ -import number_theory.legendre_symbol.quadratic_char +import number_theory.legendre_symbol.basic +import number_theory.legendre_symbol.quadratic_char.gauss_sum /-! -# Legendre symbol and quadratic reciprocity. - -This file contains results about quadratic residues modulo a prime number. - -We define the Legendre symbol $\Bigl(\frac{a}{p}\Bigr)$ as `legendre_sym p a`. -Note the order of arguments! The advantage of this form is that then `legendre_sym p` -is a multiplicative map. - -The Legendre symbol is used to define the Jacobi symbol, `jacobi_sym a b`, for integers `a` -and (odd) natural numbers `b`, which extends the Legendre symbol. +# Quadratic reciprocity. ## Main results @@ -34,7 +26,7 @@ We also prove the supplementary laws that give conditions for when `-1` or `2` ## Implementation notes The proofs use results for quadratic characters on arbitrary finite fields -from `number_theory.legendre_symbol.quadratic_char`, which in turn are based on +from `number_theory.legendre_symbol.quadratic_char.gauss_sum`, which in turn are based on properties of quadratic Gauss sums as provided by `number_theory.legendre_symbol.gauss_sum`. ## Tags @@ -44,266 +36,12 @@ quadratic residue, quadratic nonresidue, Legendre symbol, quadratic reciprocity open nat -section euler - -namespace zmod - -variables (p : ℕ) [fact p.prime] - -/-- Euler's Criterion: A unit `x` of `zmod p` is a square if and only if `x ^ (p / 2) = 1`. -/ -lemma euler_criterion_units (x : (zmod p)ˣ) : (∃ y : (zmod p)ˣ, y ^ 2 = x) ↔ x ^ (p / 2) = 1 := -begin - by_cases hc : p = 2, - { substI hc, - simp only [eq_iff_true_of_subsingleton, exists_const], }, - { have h₀ := finite_field.unit_is_square_iff (by rwa ring_char_zmod_n) x, - have hs : (∃ y : (zmod p)ˣ, y ^ 2 = x) ↔ is_square(x) := - by { rw is_square_iff_exists_sq x, - simp_rw eq_comm, }, - rw hs, - rwa card p at h₀, }, -end - -/-- Euler's Criterion: a nonzero `a : zmod p` is a square if and only if `x ^ (p / 2) = 1`. -/ -lemma euler_criterion {a : zmod p} (ha : a ≠ 0) : is_square (a : zmod p) ↔ a ^ (p / 2) = 1 := -begin - apply (iff_congr _ (by simp [units.ext_iff])).mp (euler_criterion_units p (units.mk0 a ha)), - simp only [units.ext_iff, sq, units.coe_mk0, units.coe_mul], - split, { rintro ⟨y, hy⟩, exact ⟨y, hy.symm⟩ }, - { rintro ⟨y, rfl⟩, - have hy : y ≠ 0, { rintro rfl, simpa [zero_pow] using ha, }, - refine ⟨units.mk0 y hy, _⟩, simp, } -end - -/-- If `a : zmod p` is nonzero, then `a^(p/2)` is either `1` or `-1`. -/ -lemma pow_div_two_eq_neg_one_or_one {a : zmod p} (ha : a ≠ 0) : - a ^ (p / 2) = 1 ∨ a ^ (p / 2) = -1 := -begin - cases prime.eq_two_or_odd (fact.out p.prime) with hp2 hp_odd, - { substI p, revert a ha, dec_trivial }, - rw [← mul_self_eq_one_iff, ← pow_add, ← two_mul, two_mul_odd_div_two hp_odd], - exact pow_card_sub_one_eq_one ha -end - -end zmod - -end euler - -section legendre - -/-! -### Definition of the Legendre symbol and basic properties --/ - -open zmod - -variables (p : ℕ) [fact p.prime] - -/-- The Legendre symbol of `a : ℤ` and a prime `p`, `legendre_sym p a`, -is an integer defined as - -* `0` if `a` is `0` modulo `p`; -* `1` if `a` is a nonzero square modulo `p` -* `-1` otherwise. - -Note the order of the arguments! The advantage of the order chosen here is -that `legendre_sym p` is a multiplicative function `ℤ → ℤ`. --/ -def legendre_sym (a : ℤ) : ℤ := quadratic_char (zmod p) a - -namespace legendre_sym - -/-- We have the congruence `legendre_sym p a ≡ a ^ (p / 2) mod p`. -/ -lemma eq_pow (a : ℤ) : (legendre_sym p a : zmod p) = a ^ (p / 2) := -begin - cases eq_or_ne (ring_char (zmod p)) 2 with hc hc, - { by_cases ha : (a : zmod p) = 0, - { rw [legendre_sym, ha, quadratic_char_zero, - zero_pow (nat.div_pos (fact.out p.prime).two_le (succ_pos 1))], - norm_cast, }, - { have := (ring_char_zmod_n p).symm.trans hc, -- p = 2 - substI p, - rw [legendre_sym, quadratic_char_eq_one_of_char_two hc ha], - revert ha, - generalize : (a : zmod 2) = b, revert b, dec_trivial } }, - { convert quadratic_char_eq_pow_of_char_ne_two' hc (a : zmod p), - exact (card p).symm }, -end - -/-- If `p ∤ a`, then `legendre_sym p a` is `1` or `-1`. -/ -lemma eq_one_or_neg_one {a : ℤ} (ha : (a : zmod p) ≠ 0) : - legendre_sym p a = 1 ∨ legendre_sym p a = -1 := -quadratic_char_dichotomy ha - -lemma 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 eq_zero_iff (a : ℤ) : legendre_sym p a = 0 ↔ (a : zmod p) = 0 := -quadratic_char_eq_zero_iff - -@[simp] lemma at_zero : legendre_sym p 0 = 0 := -by rw [legendre_sym, int.cast_zero, mul_char.map_zero] - -@[simp] lemma at_one : legendre_sym p 1 = 1 := -by rw [legendre_sym, int.cast_one, mul_char.map_one] - -/-- The Legendre symbol is multiplicative in `a` for `p` fixed. -/ -protected -lemma mul (a b : ℤ) : legendre_sym p (a * b) = legendre_sym p a * legendre_sym p b := -by simp only [legendre_sym, int.cast_mul, map_mul] - -/-- The Legendre symbol is a homomorphism of monoids with zero. -/ -@[simps] def hom : ℤ →*₀ ℤ := -{ to_fun := legendre_sym p, - map_zero' := at_zero p, - map_one' := at_one p, - map_mul' := legendre_sym.mul p } - -/-- The square of the symbol is 1 if `p ∤ a`. -/ -theorem sq_one {a : ℤ} (ha : (a : zmod p) ≠ 0) : (legendre_sym p a) ^ 2 = 1 := -quadratic_char_sq_one ha - -/-- The Legendre symbol of `a^2` at `p` is 1 if `p ∤ a`. -/ -theorem sq_one' {a : ℤ} (ha : (a : zmod p) ≠ 0) : legendre_sym p (a ^ 2) = 1 := -by exact_mod_cast quadratic_char_sq_one' ha - -/-- The Legendre symbol depends only on `a` mod `p`. -/ -protected -theorem mod (a : ℤ) : legendre_sym p a = legendre_sym p (a % p) := -by simp only [legendre_sym, int_cast_mod] - -/-- When `p ∤ a`, then `legendre_sym p a = 1` iff `a` is a square mod `p`. -/ -lemma eq_one_iff {a : ℤ} (ha0 : (a : zmod p) ≠ 0) : - legendre_sym p a = 1 ↔ is_square (a : zmod p) := -quadratic_char_one_iff_is_square ha0 - -lemma eq_one_iff' {a : ℕ} (ha0 : (a : zmod p) ≠ 0) : - legendre_sym p a = 1 ↔ is_square (a : zmod p) := -by {rw eq_one_iff, norm_cast, exact_mod_cast ha0} - -/-- `legendre_sym p a = -1` iff `a` is a nonsquare mod `p`. -/ -lemma eq_neg_one_iff {a : ℤ} : legendre_sym p a = -1 ↔ ¬ is_square (a : zmod p) := -quadratic_char_neg_one_iff_not_is_square - -lemma eq_neg_one_iff' {a : ℕ} : legendre_sym p a = -1 ↔ ¬ is_square (a : zmod p) := -by {rw eq_neg_one_iff, norm_cast} - -/-- The number of square roots of `a` modulo `p` is determined by the Legendre symbol. -/ -lemma card_sqrts (hp : p ≠ 2) (a : ℤ) : - ↑{x : zmod p | x^2 = a}.to_finset.card = legendre_sym p a + 1 := -quadratic_char_card_sqrts ((ring_char_zmod_n p).substr hp) a - -end legendre_sym - -end legendre - -section quadratic_form - -/-! -### Applications to binary quadratic forms --/ - -namespace legendre_sym - -/-- The Legendre symbol `legendre_sym p a = 1` if there is a solution in `ℤ/pℤ` -of the equation `x^2 - a*y^2 = 0` with `y ≠ 0`. -/ -lemma eq_one_of_sq_sub_mul_sq_eq_zero {p : ℕ} [fact p.prime] - {a : ℤ} (ha : (a : zmod p) ≠ 0) {x y : zmod p} (hy : y ≠ 0) (hxy : x ^ 2 - a * y ^ 2 = 0) : - legendre_sym p a = 1 := -begin - apply_fun (* y⁻¹ ^ 2) at hxy, - simp only [zero_mul] at hxy, - rw [(by ring : (x ^ 2 - ↑a * y ^ 2) * y⁻¹ ^ 2 = (x * y⁻¹) ^ 2 - a * (y * y⁻¹) ^ 2), - mul_inv_cancel hy, one_pow, mul_one, sub_eq_zero, pow_two] at hxy, - exact (eq_one_iff p ha).mpr ⟨x * y⁻¹, hxy.symm⟩, -end - -/-- The Legendre symbol `legendre_sym p a = 1` if there is a solution in `ℤ/pℤ` -of the equation `x^2 - a*y^2 = 0` with `x ≠ 0`. -/ -lemma eq_one_of_sq_sub_mul_sq_eq_zero' {p : ℕ} [fact p.prime] - {a : ℤ} (ha : (a : zmod p) ≠ 0) {x y : zmod p} (hx : x ≠ 0) (hxy : x ^ 2 - a * y ^ 2 = 0) : - legendre_sym p a = 1 := -begin - have hy : y ≠ 0, - { rintro rfl, - rw [zero_pow' 2 (by norm_num), mul_zero, sub_zero, pow_eq_zero_iff (by norm_num : 0 < 2)] - at hxy, - exacts [hx hxy, infer_instance], }, -- why is the instance not inferred automatically? - exact eq_one_of_sq_sub_mul_sq_eq_zero ha hy hxy, -end - -/-- If `legendre_sym p a = -1`, then the only solution of `x^2 - a*y^2 = 0` in `ℤ/pℤ` -is the trivial one. -/ -lemma eq_zero_mod_of_eq_neg_one {p : ℕ} [fact p.prime] {a : ℤ} - (h : legendre_sym p a = -1) {x y : zmod p} (hxy : x ^ 2 - a * y ^ 2 = 0) : x = 0 ∧ y = 0 := -begin - have ha : (a : zmod p) ≠ 0, - { intro hf, - rw (eq_zero_iff p a).mpr hf at h, - exact int.zero_ne_neg_of_ne zero_ne_one h, }, - by_contra hf, - cases not_and_distrib.mp hf with hx hy, - { rw [eq_one_of_sq_sub_mul_sq_eq_zero' ha hx hxy, eq_neg_self_iff] at h, - exact one_ne_zero h, }, - { rw [eq_one_of_sq_sub_mul_sq_eq_zero ha hy hxy, eq_neg_self_iff] at h, - exact one_ne_zero h, } -end - -/-- If `legendre_sym p a = -1` and `p` divides `x^2 - a*y^2`, then `p` must divide `x` and `y`. -/ -lemma prime_dvd_of_eq_neg_one {p : ℕ} [fact p.prime] {a : ℤ} - (h : legendre_sym p a = -1) {x y : ℤ} (hxy : ↑p ∣ x ^ 2 - a * y ^ 2) : ↑p ∣ x ∧ ↑p ∣ y := -begin - simp_rw ← zmod.int_coe_zmod_eq_zero_iff_dvd at hxy ⊢, - push_cast at hxy, - exact eq_zero_mod_of_eq_neg_one h hxy, -end - -end legendre_sym - -end quadratic_form - section values -/-! -### The value of the Legendre symbol at `-1` - -See `jacobi_sym.at_neg_one` for the corresponding statement for the Jacobi symbol. --/ - variables {p : ℕ} [fact p.prime] open zmod -/-- `legendre_sym p (-1)` is given by `χ₄ p`. -/ -lemma legendre_sym.at_neg_one (hp : p ≠ 2) : legendre_sym p (-1) = χ₄ p := -by simp only [legendre_sym, card p, quadratic_char_neg_one ((ring_char_zmod_n p).substr hp), - int.cast_neg, int.cast_one] - -namespace zmod - -/-- `-1` is a square in `zmod p` iff `p` is not congruent to `3` mod `4`. -/ -lemma exists_sq_eq_neg_one_iff : is_square (-1 : zmod p) ↔ p % 4 ≠ 3 := -by rw [finite_field.is_square_neg_one_iff, card p] - -lemma mod_four_ne_three_of_sq_eq_neg_one {y : zmod p} (hy : y ^ 2 = -1) : p % 4 ≠ 3 := -exists_sq_eq_neg_one_iff.1 ⟨y, hy ▸ pow_two y⟩ - -/-- If two nonzero squares are negatives of each other in `zmod p`, then `p % 4 ≠ 3`. -/ -lemma mod_four_ne_three_of_sq_eq_neg_sq' {x y : zmod p} (hy : y ≠ 0) (hxy : x ^ 2 = - y ^ 2) : - p % 4 ≠ 3 := -@mod_four_ne_three_of_sq_eq_neg_one p _ (x / y) begin - apply_fun (λ z, z / y ^ 2) at hxy, - rwa [neg_div, ←div_pow, ←div_pow, div_self hy, one_pow] at hxy -end - -lemma mod_four_ne_three_of_sq_eq_neg_sq {x y : zmod p} (hx : x ≠ 0) (hxy : x ^ 2 = - y ^ 2) : - p % 4 ≠ 3 := -mod_four_ne_three_of_sq_eq_neg_sq' hx (neg_eq_iff_eq_neg.mpr hxy).symm - -end zmod - /-! ### The value of the Legendre symbol at `2` and `-2` diff --git a/src/number_theory/sum_two_squares.lean b/src/number_theory/sum_two_squares.lean index f949ca500609d..fe7bbf6eae2f2 100644 --- a/src/number_theory/sum_two_squares.lean +++ b/src/number_theory/sum_two_squares.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes, Michael Stoll -/ -import number_theory.zsqrtd.gaussian_int +import number_theory.zsqrtd.quadratic_reciprocity import tactic.linear_combination /-! diff --git a/src/number_theory/zsqrtd/gaussian_int.lean b/src/number_theory/zsqrtd/gaussian_int.lean index 66a9a7c1451b3..4c121a2f9f4d5 100644 --- a/src/number_theory/zsqrtd/gaussian_int.lean +++ b/src/number_theory/zsqrtd/gaussian_int.lean @@ -6,7 +6,8 @@ Authors: Chris Hughes import number_theory.zsqrtd.basic import data.complex.basic import ring_theory.principal_ideal_domain -import number_theory.legendre_symbol.quadratic_reciprocity + + /-! # Gaussian integers @@ -19,10 +20,11 @@ The Euclidean domain structure on `ℤ[i]` is defined in this file. The homomorphism `to_complex` into the complex numbers is also defined in this file. -## Main statements +## See also -`prime_iff_mod_four_eq_three_of_nat_prime` -A prime natural number is prime in `ℤ[i]` if and only if it is `3` mod `4` +See `number_theory.zsqrtd.gaussian_int` for: +* `prime_iff_mod_four_eq_three_of_nat_prime`: + A prime natural number is prime in `ℤ[i]` if and only if it is `3` mod `4` ## Notations @@ -194,62 +196,6 @@ instance : euclidean_domain ℤ[i] := open principal_ideal_ring -lemma mod_four_eq_three_of_nat_prime_of_prime (p : ℕ) [hp : fact p.prime] (hpi : prime (p : ℤ[i])) : - p % 4 = 3 := -hp.1.eq_two_or_odd.elim - (λ hp2, absurd hpi (mt irreducible_iff_prime.2 $ - λ ⟨hu, h⟩, begin - have := h ⟨1, 1⟩ ⟨1, -1⟩ (hp2.symm ▸ rfl), - rw [← norm_eq_one_iff, ← norm_eq_one_iff] at this, - exact absurd this dec_trivial - end)) - (λ hp1, by_contradiction $ λ hp3 : p % 4 ≠ 3, - have hp41 : p % 4 = 1, - begin - rw [← nat.mod_mul_left_mod p 2 2, show 2 * 2 = 4, from rfl] at hp1, - have := nat.mod_lt p (show 0 < 4, from dec_trivial), - revert this hp3 hp1, - generalize : p % 4 = m, dec_trivial!, - end, - let ⟨k, hk⟩ := zmod.exists_sq_eq_neg_one_iff.2 $ - by rw hp41; exact dec_trivial in - begin - obtain ⟨k, k_lt_p, rfl⟩ : ∃ (k' : ℕ) (h : k' < p), (k' : zmod p) = k, - { refine ⟨k.val, k.val_lt, zmod.nat_cast_zmod_val k⟩ }, - have hpk : p ∣ k ^ 2 + 1, - by { rw [pow_two, ← char_p.cast_eq_zero_iff (zmod p) p, nat.cast_add, nat.cast_mul, - nat.cast_one, ← hk, add_left_neg], }, - have hkmul : (k ^ 2 + 1 : ℤ[i]) = ⟨k, 1⟩ * ⟨k, -1⟩ := - by simp [sq, zsqrtd.ext], - have hpne1 : p ≠ 1 := ne_of_gt hp.1.one_lt, - have hkltp : 1 + k * k < p * p, - from calc 1 + k * k ≤ k + k * k : - add_le_add_right (nat.pos_of_ne_zero - (λ hk0, by clear_aux_decl; simp [*, pow_succ'] at *)) _ - ... = k * (k + 1) : by simp [add_comm, mul_add] - ... < p * p : mul_lt_mul k_lt_p k_lt_p (nat.succ_pos _) (nat.zero_le _), - have hpk₁ : ¬ (p : ℤ[i]) ∣ ⟨k, -1⟩ := - λ ⟨x, hx⟩, lt_irrefl (p * x : ℤ[i]).norm.nat_abs $ - calc (norm (p * x : ℤ[i])).nat_abs = (zsqrtd.norm ⟨k, -1⟩).nat_abs : by rw hx - ... < (norm (p : ℤ[i])).nat_abs : by simpa [add_comm, zsqrtd.norm] using hkltp - ... ≤ (norm (p * x : ℤ[i])).nat_abs : norm_le_norm_mul_left _ - (λ hx0, (show (-1 : ℤ) ≠ 0, from dec_trivial) $ - by simpa [hx0] using congr_arg zsqrtd.im hx), - have hpk₂ : ¬ (p : ℤ[i]) ∣ ⟨k, 1⟩ := - λ ⟨x, hx⟩, lt_irrefl (p * x : ℤ[i]).norm.nat_abs $ - calc (norm (p * x : ℤ[i])).nat_abs = (zsqrtd.norm ⟨k, 1⟩).nat_abs : by rw hx - ... < (norm (p : ℤ[i])).nat_abs : by simpa [add_comm, zsqrtd.norm] using hkltp - ... ≤ (norm (p * x : ℤ[i])).nat_abs : norm_le_norm_mul_left _ - (λ hx0, (show (1 : ℤ) ≠ 0, from dec_trivial) $ - by simpa [hx0] using congr_arg zsqrtd.im hx), - have hpu : ¬ is_unit (p : ℤ[i]), from mt norm_eq_one_iff.2 - (by rw [norm_nat_cast, int.nat_abs_mul, mul_eq_one]; - exact λ h, (ne_of_lt hp.1.one_lt).symm h.1), - obtain ⟨y, hy⟩ := hpk, - have := hpi.2.2 ⟨k, 1⟩ ⟨k, -1⟩ ⟨y, by rw [← hkmul, ← nat.cast_mul p, ← hy]; simp⟩, - clear_aux_decl, tauto - end) - lemma sq_add_sq_of_nat_prime_of_not_irreducible (p : ℕ) [hp : fact p.prime] (hpi : ¬irreducible (p : ℤ[i])) : ∃ a b, a^2 + b^2 = p := have hpu : ¬ is_unit (p : ℤ[i]), from mt norm_eq_one_iff.2 $ @@ -265,16 +211,4 @@ have hnap : (norm a).nat_abs = p, from ((hp.1.mul_eq_prime_sq_iff simp).1, ⟨a.re.nat_abs, a.im.nat_abs, by simpa [nat_abs_norm_eq, sq] using hnap⟩ -lemma prime_of_nat_prime_of_mod_four_eq_three (p : ℕ) [hp : fact p.prime] (hp3 : p % 4 = 3) : - prime (p : ℤ[i]) := -irreducible_iff_prime.1 $ classical.by_contradiction $ λ hpi, - let ⟨a, b, hab⟩ := sq_add_sq_of_nat_prime_of_not_irreducible p hpi in -have ∀ a b : zmod 4, a^2 + b^2 ≠ p, by erw [← zmod.nat_cast_mod p 4, hp3]; exact dec_trivial, -this a b (hab ▸ by simp) - -/-- A prime natural number is prime in `ℤ[i]` if and only if it is `3` mod `4` -/ -lemma prime_iff_mod_four_eq_three_of_nat_prime (p : ℕ) [hp : fact p.prime] : - prime (p : ℤ[i]) ↔ p % 4 = 3 := -⟨mod_four_eq_three_of_nat_prime_of_prime p, prime_of_nat_prime_of_mod_four_eq_three p⟩ - end gaussian_int diff --git a/src/number_theory/zsqrtd/quadratic_reciprocity.lean b/src/number_theory/zsqrtd/quadratic_reciprocity.lean new file mode 100644 index 0000000000000..876937bec856e --- /dev/null +++ b/src/number_theory/zsqrtd/quadratic_reciprocity.lean @@ -0,0 +1,96 @@ +/- +Copyright (c) 2019 Chris Hughes. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Chris Hughes +-/ +import number_theory.zsqrtd.gaussian_int +import number_theory.legendre_symbol.quadratic_reciprocity + +/-! +# Facts about the gaussian integers relying on quadratic reciprocity. + +## Main statements + +`prime_iff_mod_four_eq_three_of_nat_prime` +A prime natural number is prime in `ℤ[i]` if and only if it is `3` mod `4` + +-/ + +open zsqrtd complex +open_locale complex_conjugate + +local notation `ℤ[i]` := gaussian_int + +namespace gaussian_int + +open principal_ideal_ring + +lemma mod_four_eq_three_of_nat_prime_of_prime (p : ℕ) [hp : fact p.prime] (hpi : prime (p : ℤ[i])) : + p % 4 = 3 := +hp.1.eq_two_or_odd.elim + (λ hp2, absurd hpi (mt irreducible_iff_prime.2 $ + λ ⟨hu, h⟩, begin + have := h ⟨1, 1⟩ ⟨1, -1⟩ (hp2.symm ▸ rfl), + rw [← norm_eq_one_iff, ← norm_eq_one_iff] at this, + exact absurd this dec_trivial + end)) + (λ hp1, by_contradiction $ λ hp3 : p % 4 ≠ 3, + have hp41 : p % 4 = 1, + begin + rw [← nat.mod_mul_left_mod p 2 2, show 2 * 2 = 4, from rfl] at hp1, + have := nat.mod_lt p (show 0 < 4, from dec_trivial), + revert this hp3 hp1, + generalize : p % 4 = m, dec_trivial!, + end, + let ⟨k, hk⟩ := zmod.exists_sq_eq_neg_one_iff.2 $ + by rw hp41; exact dec_trivial in + begin + obtain ⟨k, k_lt_p, rfl⟩ : ∃ (k' : ℕ) (h : k' < p), (k' : zmod p) = k, + { refine ⟨k.val, k.val_lt, zmod.nat_cast_zmod_val k⟩ }, + have hpk : p ∣ k ^ 2 + 1, + by { rw [pow_two, ← char_p.cast_eq_zero_iff (zmod p) p, nat.cast_add, nat.cast_mul, + nat.cast_one, ← hk, add_left_neg], }, + have hkmul : (k ^ 2 + 1 : ℤ[i]) = ⟨k, 1⟩ * ⟨k, -1⟩ := + by simp [sq, zsqrtd.ext], + have hpne1 : p ≠ 1 := ne_of_gt hp.1.one_lt, + have hkltp : 1 + k * k < p * p, + from calc 1 + k * k ≤ k + k * k : + add_le_add_right (nat.pos_of_ne_zero + (λ hk0, by clear_aux_decl; simp [*, pow_succ'] at *)) _ + ... = k * (k + 1) : by simp [add_comm, mul_add] + ... < p * p : mul_lt_mul k_lt_p k_lt_p (nat.succ_pos _) (nat.zero_le _), + have hpk₁ : ¬ (p : ℤ[i]) ∣ ⟨k, -1⟩ := + λ ⟨x, hx⟩, lt_irrefl (p * x : ℤ[i]).norm.nat_abs $ + calc (norm (p * x : ℤ[i])).nat_abs = (zsqrtd.norm ⟨k, -1⟩).nat_abs : by rw hx + ... < (norm (p : ℤ[i])).nat_abs : by simpa [add_comm, zsqrtd.norm] using hkltp + ... ≤ (norm (p * x : ℤ[i])).nat_abs : norm_le_norm_mul_left _ + (λ hx0, (show (-1 : ℤ) ≠ 0, from dec_trivial) $ + by simpa [hx0] using congr_arg zsqrtd.im hx), + have hpk₂ : ¬ (p : ℤ[i]) ∣ ⟨k, 1⟩ := + λ ⟨x, hx⟩, lt_irrefl (p * x : ℤ[i]).norm.nat_abs $ + calc (norm (p * x : ℤ[i])).nat_abs = (zsqrtd.norm ⟨k, 1⟩).nat_abs : by rw hx + ... < (norm (p : ℤ[i])).nat_abs : by simpa [add_comm, zsqrtd.norm] using hkltp + ... ≤ (norm (p * x : ℤ[i])).nat_abs : norm_le_norm_mul_left _ + (λ hx0, (show (1 : ℤ) ≠ 0, from dec_trivial) $ + by simpa [hx0] using congr_arg zsqrtd.im hx), + have hpu : ¬ is_unit (p : ℤ[i]), from mt norm_eq_one_iff.2 + (by rw [norm_nat_cast, int.nat_abs_mul, mul_eq_one]; + exact λ h, (ne_of_lt hp.1.one_lt).symm h.1), + obtain ⟨y, hy⟩ := hpk, + have := hpi.2.2 ⟨k, 1⟩ ⟨k, -1⟩ ⟨y, by rw [← hkmul, ← nat.cast_mul p, ← hy]; simp⟩, + clear_aux_decl, tauto + end) + +lemma prime_of_nat_prime_of_mod_four_eq_three (p : ℕ) [hp : fact p.prime] (hp3 : p % 4 = 3) : + prime (p : ℤ[i]) := +irreducible_iff_prime.1 $ classical.by_contradiction $ λ hpi, + let ⟨a, b, hab⟩ := sq_add_sq_of_nat_prime_of_not_irreducible p hpi in +have ∀ a b : zmod 4, a^2 + b^2 ≠ p, by erw [← zmod.nat_cast_mod p 4, hp3]; exact dec_trivial, +this a b (hab ▸ by simp) + +/-- A prime natural number is prime in `ℤ[i]` if and only if it is `3` mod `4` -/ +lemma prime_iff_mod_four_eq_three_of_nat_prime (p : ℕ) [hp : fact p.prime] : + prime (p : ℤ[i]) ↔ p % 4 = 3 := +⟨mod_four_eq_three_of_nat_prime_of_prime p, prime_of_nat_prime_of_mod_four_eq_three p⟩ + +end gaussian_int From cff8231f04dfa33fd8f2f45792eebd862ef30cad Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Fri, 16 Jun 2023 05:28:54 +0000 Subject: [PATCH 1194/1291] chore(*): add mathlib4 synchronization comments (#19194) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.continued_fractions.computation.approximations` * `analysis.inner_product_space.rayleigh` * `analysis.inner_product_space.spectrum` * `analysis.normed.group.SemiNormedGroup` * `category_theory.closed.functor` * `field_theory.splitting_field.construction` * `geometry.euclidean.angle.sphere` * `linear_algebra.matrix.pos_def` * `linear_algebra.matrix.spectrum` * `order.category.BoolAlg` * `order.category.FinBddDistLat` * `ring_theory.local_properties` * `ring_theory.ring_hom.surjective` * `topology.sheaves.locally_surjective` --- .../continued_fractions/computation/approximations.lean | 3 +++ src/analysis/inner_product_space/rayleigh.lean | 3 +++ src/analysis/inner_product_space/spectrum.lean | 3 +++ src/analysis/normed/group/SemiNormedGroup.lean | 3 +++ src/category_theory/closed/functor.lean | 3 +++ src/field_theory/splitting_field/construction.lean | 3 +++ src/geometry/euclidean/angle/sphere.lean | 3 +++ src/linear_algebra/matrix/pos_def.lean | 3 +++ src/linear_algebra/matrix/spectrum.lean | 3 +++ src/order/category/BoolAlg.lean | 3 +++ src/order/category/FinBddDistLat.lean | 3 +++ src/ring_theory/local_properties.lean | 3 +++ src/ring_theory/ring_hom/surjective.lean | 3 +++ src/topology/sheaves/locally_surjective.lean | 3 +++ 14 files changed, 42 insertions(+) diff --git a/src/algebra/continued_fractions/computation/approximations.lean b/src/algebra/continued_fractions/computation/approximations.lean index 19ca7909b95c2..09633aa94d4c2 100644 --- a/src/algebra/continued_fractions/computation/approximations.lean +++ b/src/algebra/continued_fractions/computation/approximations.lean @@ -9,6 +9,9 @@ import tactic.solve_by_elim /-! # Approximations for Continued Fraction Computations (`generalized_continued_fraction.of`) +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Summary This file contains useful approximations for the values involved in the continued fractions diff --git a/src/analysis/inner_product_space/rayleigh.lean b/src/analysis/inner_product_space/rayleigh.lean index a01efa5925ee3..7d9b1bed9dbca 100644 --- a/src/analysis/inner_product_space/rayleigh.lean +++ b/src/analysis/inner_product_space/rayleigh.lean @@ -12,6 +12,9 @@ import linear_algebra.eigenspace.basic /-! # The Rayleigh quotient +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The Rayleigh quotient of a self-adjoint operator `T` on an inner product space `E` is the function `λ x, ⟪T x, x⟫ / ‖x‖ ^ 2`. diff --git a/src/analysis/inner_product_space/spectrum.lean b/src/analysis/inner_product_space/spectrum.lean index d3a9f2866d8e8..6c3e285e142f5 100644 --- a/src/analysis/inner_product_space/spectrum.lean +++ b/src/analysis/inner_product_space/spectrum.lean @@ -10,6 +10,9 @@ import linear_algebra.eigenspace.minpoly /-! # Spectral theory of self-adjoint operators +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file covers the spectral theory of self-adjoint operators on an inner product space. The first part of the file covers general properties, true without any condition on boundedness or diff --git a/src/analysis/normed/group/SemiNormedGroup.lean b/src/analysis/normed/group/SemiNormedGroup.lean index c06355e1fe803..db67750524950 100644 --- a/src/analysis/normed/group/SemiNormedGroup.lean +++ b/src/analysis/normed/group/SemiNormedGroup.lean @@ -11,6 +11,9 @@ import category_theory.elementwise /-! # The category of seminormed groups +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define `SemiNormedGroup`, the category of seminormed groups and normed group homs between them, as well as `SemiNormedGroup₁`, the subcategory of norm non-increasing morphisms. -/ diff --git a/src/category_theory/closed/functor.lean b/src/category_theory/closed/functor.lean index 44589883772c8..ef432168151b0 100644 --- a/src/category_theory/closed/functor.lean +++ b/src/category_theory/closed/functor.lean @@ -11,6 +11,9 @@ import category_theory.adjunction.fully_faithful /-! # Cartesian closed functors +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Define the exponential comparison morphisms for a functor which preserves binary products, and use them to define a cartesian closed functor: one which (naturally) preserves exponentials. diff --git a/src/field_theory/splitting_field/construction.lean b/src/field_theory/splitting_field/construction.lean index 6c9a641450ca0..85ec974bcb001 100644 --- a/src/field_theory/splitting_field/construction.lean +++ b/src/field_theory/splitting_field/construction.lean @@ -8,6 +8,9 @@ import field_theory.normal /-! # Splitting fields +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove the existence and uniqueness of splitting fields. ## Main definitions diff --git a/src/geometry/euclidean/angle/sphere.lean b/src/geometry/euclidean/angle/sphere.lean index 00f153eef699e..4b32efd19ef3a 100644 --- a/src/geometry/euclidean/angle/sphere.lean +++ b/src/geometry/euclidean/angle/sphere.lean @@ -9,6 +9,9 @@ import geometry.euclidean.circumcenter /-! # Angles in circles and sphere. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves results about angles in circles and spheres. -/ diff --git a/src/linear_algebra/matrix/pos_def.lean b/src/linear_algebra/matrix/pos_def.lean index 6e224830d8571..8a06f1f566eb2 100644 --- a/src/linear_algebra/matrix/pos_def.lean +++ b/src/linear_algebra/matrix/pos_def.lean @@ -7,6 +7,9 @@ import linear_algebra.matrix.spectrum import linear_algebra.quadratic_form.basic /-! # Positive Definite Matrices + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. This file defines positive (semi)definite matrices and connects the notion to positive definiteness of quadratic forms. ## Main definition diff --git a/src/linear_algebra/matrix/spectrum.lean b/src/linear_algebra/matrix/spectrum.lean index 0d533048e00fe..e7f85b320877c 100644 --- a/src/linear_algebra/matrix/spectrum.lean +++ b/src/linear_algebra/matrix/spectrum.lean @@ -8,6 +8,9 @@ import linear_algebra.matrix.hermitian /-! # Spectral theory of hermitian matrices +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves the spectral theorem for matrices. The proof of the spectral theorem is based on the spectral theorem for linear maps (`diagonalization_basis_apply_self_apply`). diff --git a/src/order/category/BoolAlg.lean b/src/order/category/BoolAlg.lean index 08e38c978beb0..a84443001449e 100644 --- a/src/order/category/BoolAlg.lean +++ b/src/order/category/BoolAlg.lean @@ -8,6 +8,9 @@ import order.category.HeytAlg /-! # The category of boolean algebras +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This defines `BoolAlg`, the category of boolean algebras. -/ diff --git a/src/order/category/FinBddDistLat.lean b/src/order/category/FinBddDistLat.lean index 2e04a3deab863..02344d9e5de74 100644 --- a/src/order/category/FinBddDistLat.lean +++ b/src/order/category/FinBddDistLat.lean @@ -10,6 +10,9 @@ import order.category.FinPartOrd /-! # The category of finite bounded distributive lattices +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines `FinBddDistLat`, the category of finite distributive lattices with bounded lattice homomorphisms. -/ diff --git a/src/ring_theory/local_properties.lean b/src/ring_theory/local_properties.lean index c958abad5cea5..47bf634d18d70 100644 --- a/src/ring_theory/local_properties.lean +++ b/src/ring_theory/local_properties.lean @@ -14,6 +14,9 @@ import ring_theory.ring_hom_properties /-! # Local properties of commutative rings +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we provide the proofs of various local properties. ## Naming Conventions diff --git a/src/ring_theory/ring_hom/surjective.lean b/src/ring_theory/ring_hom/surjective.lean index ae8a364937991..82cf95993fdab 100644 --- a/src/ring_theory/ring_hom/surjective.lean +++ b/src/ring_theory/ring_hom/surjective.lean @@ -9,6 +9,9 @@ import ring_theory.local_properties # The meta properties of surjective ring homomorphisms. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + -/ namespace ring_hom diff --git a/src/topology/sheaves/locally_surjective.lean b/src/topology/sheaves/locally_surjective.lean index a1c6a63df1dd8..c71ef4918c9d0 100644 --- a/src/topology/sheaves/locally_surjective.lean +++ b/src/topology/sheaves/locally_surjective.lean @@ -12,6 +12,9 @@ import category_theory.sites.surjective # Locally surjective maps of presheaves. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Let `X` be a topological space, `ℱ` and `𝒢` presheaves on `X`, `T : ℱ ⟶ 𝒢` a map. In this file we formulate two notions for what it means for From 5563b1b49e86e135e8c7b556da5ad2f5ff881cad Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Pol=27tta=20/=20Miyahara=20K=C5=8D?= Date: Fri, 16 Jun 2023 05:28:55 +0000 Subject: [PATCH 1195/1291] style(archive.100-theorems-list.xx_theorem_name): rename to `wiedijk_100_theorems.theorem_name` (#19195) https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/Archive.2E100-theorems-list.2EXXTheoremName.20causes.20error.20!4.235114 --- archive/{100-theorems-list => wiedijk_100_theorems}/README.md | 0 .../abel_ruffini.lean} | 0 .../area_of_a_circle.lean} | 0 .../ascending_descending_sequences.lean} | 0 .../ballot_problem.lean} | 0 .../birthday_problem.lean} | 0 .../cubing_a_cube.lean} | 0 .../friendship_graphs.lean} | 0 .../herons_formula.lean} | 0 .../inverse_triangle_sum.lean} | 0 .../54_konigsberg.lean => wiedijk_100_theorems/konigsberg.lean} | 0 .../45_partition.lean => wiedijk_100_theorems/partition.lean} | 0 .../perfect_numbers.lean} | 0 .../solution_of_cubic.lean} | 0 .../sum_of_prime_reciprocals_diverges.lean} | 0 15 files changed, 0 insertions(+), 0 deletions(-) rename archive/{100-theorems-list => wiedijk_100_theorems}/README.md (100%) rename archive/{100-theorems-list/16_abel_ruffini.lean => wiedijk_100_theorems/abel_ruffini.lean} (100%) rename archive/{100-theorems-list/9_area_of_a_circle.lean => wiedijk_100_theorems/area_of_a_circle.lean} (100%) rename archive/{100-theorems-list/73_ascending_descending_sequences.lean => wiedijk_100_theorems/ascending_descending_sequences.lean} (100%) rename archive/{100-theorems-list/30_ballot_problem.lean => wiedijk_100_theorems/ballot_problem.lean} (100%) rename archive/{100-theorems-list/93_birthday_problem.lean => wiedijk_100_theorems/birthday_problem.lean} (100%) rename archive/{100-theorems-list/82_cubing_a_cube.lean => wiedijk_100_theorems/cubing_a_cube.lean} (100%) rename archive/{100-theorems-list/83_friendship_graphs.lean => wiedijk_100_theorems/friendship_graphs.lean} (100%) rename archive/{100-theorems-list/57_herons_formula.lean => wiedijk_100_theorems/herons_formula.lean} (100%) rename archive/{100-theorems-list/42_inverse_triangle_sum.lean => wiedijk_100_theorems/inverse_triangle_sum.lean} (100%) rename archive/{100-theorems-list/54_konigsberg.lean => wiedijk_100_theorems/konigsberg.lean} (100%) rename archive/{100-theorems-list/45_partition.lean => wiedijk_100_theorems/partition.lean} (100%) rename archive/{100-theorems-list/70_perfect_numbers.lean => wiedijk_100_theorems/perfect_numbers.lean} (100%) rename archive/{100-theorems-list/37_solution_of_cubic.lean => wiedijk_100_theorems/solution_of_cubic.lean} (100%) rename archive/{100-theorems-list/81_sum_of_prime_reciprocals_diverges.lean => wiedijk_100_theorems/sum_of_prime_reciprocals_diverges.lean} (100%) diff --git a/archive/100-theorems-list/README.md b/archive/wiedijk_100_theorems/README.md similarity index 100% rename from archive/100-theorems-list/README.md rename to archive/wiedijk_100_theorems/README.md diff --git a/archive/100-theorems-list/16_abel_ruffini.lean b/archive/wiedijk_100_theorems/abel_ruffini.lean similarity index 100% rename from archive/100-theorems-list/16_abel_ruffini.lean rename to archive/wiedijk_100_theorems/abel_ruffini.lean diff --git a/archive/100-theorems-list/9_area_of_a_circle.lean b/archive/wiedijk_100_theorems/area_of_a_circle.lean similarity index 100% rename from archive/100-theorems-list/9_area_of_a_circle.lean rename to archive/wiedijk_100_theorems/area_of_a_circle.lean diff --git a/archive/100-theorems-list/73_ascending_descending_sequences.lean b/archive/wiedijk_100_theorems/ascending_descending_sequences.lean similarity index 100% rename from archive/100-theorems-list/73_ascending_descending_sequences.lean rename to archive/wiedijk_100_theorems/ascending_descending_sequences.lean diff --git a/archive/100-theorems-list/30_ballot_problem.lean b/archive/wiedijk_100_theorems/ballot_problem.lean similarity index 100% rename from archive/100-theorems-list/30_ballot_problem.lean rename to archive/wiedijk_100_theorems/ballot_problem.lean diff --git a/archive/100-theorems-list/93_birthday_problem.lean b/archive/wiedijk_100_theorems/birthday_problem.lean similarity index 100% rename from archive/100-theorems-list/93_birthday_problem.lean rename to archive/wiedijk_100_theorems/birthday_problem.lean diff --git a/archive/100-theorems-list/82_cubing_a_cube.lean b/archive/wiedijk_100_theorems/cubing_a_cube.lean similarity index 100% rename from archive/100-theorems-list/82_cubing_a_cube.lean rename to archive/wiedijk_100_theorems/cubing_a_cube.lean diff --git a/archive/100-theorems-list/83_friendship_graphs.lean b/archive/wiedijk_100_theorems/friendship_graphs.lean similarity index 100% rename from archive/100-theorems-list/83_friendship_graphs.lean rename to archive/wiedijk_100_theorems/friendship_graphs.lean diff --git a/archive/100-theorems-list/57_herons_formula.lean b/archive/wiedijk_100_theorems/herons_formula.lean similarity index 100% rename from archive/100-theorems-list/57_herons_formula.lean rename to archive/wiedijk_100_theorems/herons_formula.lean diff --git a/archive/100-theorems-list/42_inverse_triangle_sum.lean b/archive/wiedijk_100_theorems/inverse_triangle_sum.lean similarity index 100% rename from archive/100-theorems-list/42_inverse_triangle_sum.lean rename to archive/wiedijk_100_theorems/inverse_triangle_sum.lean diff --git a/archive/100-theorems-list/54_konigsberg.lean b/archive/wiedijk_100_theorems/konigsberg.lean similarity index 100% rename from archive/100-theorems-list/54_konigsberg.lean rename to archive/wiedijk_100_theorems/konigsberg.lean diff --git a/archive/100-theorems-list/45_partition.lean b/archive/wiedijk_100_theorems/partition.lean similarity index 100% rename from archive/100-theorems-list/45_partition.lean rename to archive/wiedijk_100_theorems/partition.lean diff --git a/archive/100-theorems-list/70_perfect_numbers.lean b/archive/wiedijk_100_theorems/perfect_numbers.lean similarity index 100% rename from archive/100-theorems-list/70_perfect_numbers.lean rename to archive/wiedijk_100_theorems/perfect_numbers.lean diff --git a/archive/100-theorems-list/37_solution_of_cubic.lean b/archive/wiedijk_100_theorems/solution_of_cubic.lean similarity index 100% rename from archive/100-theorems-list/37_solution_of_cubic.lean rename to archive/wiedijk_100_theorems/solution_of_cubic.lean diff --git a/archive/100-theorems-list/81_sum_of_prime_reciprocals_diverges.lean b/archive/wiedijk_100_theorems/sum_of_prime_reciprocals_diverges.lean similarity index 100% rename from archive/100-theorems-list/81_sum_of_prime_reciprocals_diverges.lean rename to archive/wiedijk_100_theorems/sum_of_prime_reciprocals_diverges.lean From c4c2ed622f43768eff32608d4a0f8a6cec1c047d Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 16 Jun 2023 08:34:16 +0000 Subject: [PATCH 1196/1291] feat(order/sup_indep): More lemmas (#11932) A few more lemmas about `finset.sup_indep` and `set.pairwise_disjoint`. --- src/data/finset/pairwise.lean | 14 +++- src/data/set/pairwise/basic.lean | 16 ++++- src/data/set/pairwise/lattice.lean | 41 ++++++++++-- src/data/set/prod.lean | 24 ++++++- src/order/sup_indep.lean | 100 ++++++++++++++++++++++++++++- 5 files changed, 185 insertions(+), 10 deletions(-) diff --git a/src/data/finset/pairwise.lean b/src/data/finset/pairwise.lean index d71306e29b6c1..4cdbc55112e2a 100644 --- a/src/data/finset/pairwise.lean +++ b/src/data/finset/pairwise.lean @@ -39,15 +39,23 @@ lemma pairwise_disjoint.elim_finset {s : set ι} {f : ι → finset α} i = j := hs.elim hi hj (finset.not_disjoint_iff.2 ⟨a, hai, haj⟩) -lemma pairwise_disjoint.image_finset_of_le [decidable_eq ι] [semilattice_inf α] [order_bot α] - {s : finset ι} {f : ι → α} (hs : (s : set ι).pairwise_disjoint f) {g : ι → ι} - (hf : ∀ a, f (g a) ≤ f a) : +section semilattice_inf +variables [semilattice_inf α] [order_bot α] {s : finset ι} {f : ι → α} + +lemma pairwise_disjoint.image_finset_of_le [decidable_eq ι] {s : finset ι} {f : ι → α} + (hs : (s : set ι).pairwise_disjoint f) {g : ι → ι} (hf : ∀ a, f (g a) ≤ f a) : (s.image g : set ι).pairwise_disjoint f := begin rw coe_image, exact hs.image_of_le hf, end +lemma pairwise_disjoint.attach (hs : (s : set ι).pairwise_disjoint f) : + (s.attach : set {x // x ∈ s}).pairwise_disjoint (f ∘ subtype.val) := +λ i _ j _ hij, hs i.2 j.2 $ mt subtype.ext_val hij + +end semilattice_inf + variables [lattice α] [order_bot α] /-- Bind operation for `set.pairwise_disjoint`. In a complete lattice, you can use diff --git a/src/data/set/pairwise/basic.lean b/src/data/set/pairwise/basic.lean index 04aa55cfed3c7..25b2a2fac226e 100644 --- a/src/data/set/pairwise/basic.lean +++ b/src/data/set/pairwise/basic.lean @@ -30,7 +30,7 @@ The spelling `s.pairwise_disjoint id` is preferred over `s.pairwise disjoint` to on `set.pairwise_disjoint`, even though the latter unfolds to something nicer. -/ -open set function +open function order set variables {α β γ ι ι' : Type*} {r p q : α → α → Prop} @@ -296,6 +296,8 @@ end semilattice_inf_bot /-! ### Pairwise disjoint set of sets -/ +variables {s : set ι} {t : set ι'} + lemma pairwise_disjoint_range_singleton : (set.range (singleton : ι → set ι)).pairwise_disjoint id := begin @@ -311,6 +313,18 @@ lemma pairwise_disjoint.elim_set {s : set ι} {f : ι → set α} (hs : s.pairwi (hi : i ∈ s) (hj : j ∈ s) (a : α) (hai : a ∈ f i) (haj : a ∈ f j) : i = j := hs.elim hi hj $ not_disjoint_iff.2 ⟨a, hai, haj⟩ +lemma pairwise_disjoint.prod {f : ι → set α} {g : ι' → set β} (hs : s.pairwise_disjoint f) + (ht : t.pairwise_disjoint g) : + (s ×ˢ t : set (ι × ι')).pairwise_disjoint (λ i, f i.1 ×ˢ g i.2) := +λ ⟨i, i'⟩ ⟨hi, hi'⟩ ⟨j, j'⟩ ⟨hj, hj'⟩ hij, disjoint_left.2 $ λ ⟨a, b⟩ ⟨hai, hbi⟩ ⟨haj, hbj⟩, + hij $ prod.ext (hs.elim_set hi hj _ hai haj) $ ht.elim_set hi' hj' _ hbi hbj + +lemma pairwise_disjoint_pi {ι' α : ι → Type*} {s : Π i, set (ι' i)} {f : Π i, ι' i → set (α i)} + (hs : ∀ i, (s i).pairwise_disjoint (f i)) : + ((univ : set ι).pi s).pairwise_disjoint (λ I, (univ : set ι).pi (λ i, f _ (I i))) := +λ I hI J hJ hIJ, disjoint_left.2 $ λ a haI haJ, hIJ $ funext $ λ i, + (hs i).elim_set (hI i trivial) (hJ i trivial) (a i) (haI i trivial) (haJ i trivial) + /-- The partial images of a binary function `f` whose partial evaluations are injective are pairwise disjoint iff `f` is injective . -/ lemma pairwise_disjoint_image_right_iff {f : α → β → γ} {s : set α} {t : set β} diff --git a/src/data/set/pairwise/lattice.lean b/src/data/set/pairwise/lattice.lean index 1b27355c2f037..3382d20222314 100644 --- a/src/data/set/pairwise/lattice.lean +++ b/src/data/set/pairwise/lattice.lean @@ -15,16 +15,16 @@ import data.set.pairwise.basic In this file we prove many facts about `pairwise` and the set lattice. -/ -open set function +open function set order -variables {α β γ ι ι' : Type*} {r p q : α → α → Prop} +variables {α β γ ι ι' : Type*} {κ : Sort*} {r p q : α → α → Prop} section pairwise variables {f g : ι → α} {s t u : set α} {a b : α} namespace set -lemma pairwise_Union {f : ι → set α} (h : directed (⊆) f) : +lemma pairwise_Union {f : κ → set α} (h : directed (⊆) f) : (⋃ n, f n).pairwise r ↔ ∀ n, (f n).pairwise r := begin split, @@ -60,7 +60,7 @@ pairwise_sUnion h end partial_order_bot section complete_lattice -variables [complete_lattice α] +variables [complete_lattice α] {s : set ι} {t : set ι'} /-- Bind operation for `set.pairwise_disjoint`. If you want to only consider finsets of indices, you can use `set.pairwise_disjoint.bUnion_finset`. -/ @@ -78,8 +78,41 @@ begin { exact (hs hc hd $ ne_of_apply_ne _ hcd).mono (le_supr₂ a ha) (le_supr₂ b hb) } end +/-- If the suprema of columns are pairwise disjoint and suprema of rows as well, then everything is +pairwise disjoint. Not to be confused with `set.pairwise_disjoint.prod`. -/ +lemma pairwise_disjoint.prod_left {f : ι × ι' → α} + (hs : s.pairwise_disjoint $ λ i, ⨆ i' ∈ t, f (i, i')) + (ht : t.pairwise_disjoint $ λ i', ⨆ i ∈ s, f (i, i')) : + (s ×ˢ t : set (ι × ι')).pairwise_disjoint f := +begin + rintro ⟨i, i'⟩ hi ⟨j, j'⟩ hj h, + rw mem_prod at hi hj, + obtain rfl | hij := eq_or_ne i j, + { refine (ht hi.2 hj.2 $ (prod.mk.inj_left _).ne_iff.1 h).mono _ _, + { convert le_supr₂ i hi.1, refl }, + { convert le_supr₂ i hj.1, refl } }, + { refine (hs hi.1 hj.1 hij).mono _ _, + { convert le_supr₂ i' hi.2, refl }, + { convert le_supr₂ j' hj.2, refl } } +end + end complete_lattice +section frame +variables [frame α] + +lemma pairwise_disjoint_prod_left {s : set ι} {t : set ι'} {f : ι × ι' → α} : + (s ×ˢ t : set (ι × ι')).pairwise_disjoint f ↔ s.pairwise_disjoint (λ i, ⨆ i' ∈ t, f (i, i')) ∧ + t.pairwise_disjoint (λ i', ⨆ i ∈ s, f (i, i')) := +begin + refine (⟨λ h, ⟨λ i hi j hj hij, _, λ i hi j hj hij, _⟩, λ h, h.1.prod_left h.2⟩); + simp_rw [function.on_fun, supr_disjoint_iff, disjoint_supr_iff]; intros i' hi' j' hj', + { exact h (mk_mem_prod hi hi') (mk_mem_prod hj hj') (ne_of_apply_ne prod.fst hij) }, + { exact h (mk_mem_prod hi' hi) (mk_mem_prod hj' hj) (ne_of_apply_ne prod.snd hij) } +end + +end frame + lemma bUnion_diff_bUnion_eq {s t : set ι} {f : ι → set α} (h : (s ∪ t).pairwise_disjoint f) : (⋃ i ∈ s, f i) \ (⋃ i ∈ t, f i) = (⋃ i ∈ s \ t, f i) := begin diff --git a/src/data/set/prod.lean b/src/data/set/prod.lean index dc29231b9b051..e6f35bb6ad86c 100644 --- a/src/data/set/prod.lean +++ b/src/data/set/prod.lean @@ -106,13 +106,21 @@ by { ext ⟨x, y⟩, simp only [←and_and_distrib_left, mem_inter_iff, mem_prod lemma prod_inter_prod : s₁ ×ˢ t₁ ∩ s₂ ×ˢ t₂ = (s₁ ∩ s₂) ×ˢ (t₁ ∩ t₂) := by { ext ⟨x, y⟩, simp [and_assoc, and.left_comm] } -lemma disjoint_prod : disjoint (s₁ ×ˢ t₁) (s₂ ×ˢ t₂) ↔ disjoint s₁ s₂ ∨ disjoint t₁ t₂ := +@[simp] lemma disjoint_prod : disjoint (s₁ ×ˢ t₁) (s₂ ×ˢ t₂) ↔ disjoint s₁ s₂ ∨ disjoint t₁ t₂ := begin simp_rw [disjoint_left, mem_prod, not_and_distrib, prod.forall, and_imp, ←@forall_or_distrib_right α, ←@forall_or_distrib_left β, ←@forall_or_distrib_right (_ ∈ s₁), ←@forall_or_distrib_left (_ ∈ t₁)], end +lemma _root_.disjoint.set_prod_left (hs : disjoint s₁ s₂) (t₁ t₂ : set β) : + disjoint (s₁ ×ˢ t₁) (s₂ ×ˢ t₂) := +disjoint_left.2 $ λ ⟨a, b⟩ ⟨ha₁, hb₁⟩ ⟨ha₂, hb₂⟩, disjoint_left.1 hs ha₁ ha₂ + +lemma _root_.disjoint.set_prod_right (ht : disjoint t₁ t₂) (s₁ s₂ : set α) : + disjoint (s₁ ×ˢ t₁) (s₂ ×ˢ t₂) := +disjoint_left.2 $ λ ⟨a, b⟩ ⟨ha₁, hb₁⟩ ⟨ha₂, hb₂⟩, disjoint_left.1 ht hb₁ hb₂ + lemma insert_prod : insert a s ×ˢ t = (prod.mk a '' t) ∪ s ×ˢ t := by { ext ⟨x, y⟩, simp [image, iff_def, or_imp_distrib, imp.swap] {contextual := tt} } @@ -479,6 +487,20 @@ univ_pi_eq_empty_iff.2 $ h.elim $ λ x, ⟨x, rfl⟩ @[simp] lemma disjoint_univ_pi : disjoint (pi univ t₁) (pi univ t₂) ↔ ∃ i, disjoint (t₁ i) (t₂ i) := by simp only [disjoint_iff_inter_eq_empty, ← pi_inter_distrib, univ_pi_eq_empty_iff] +lemma _root_.disjoint.set_pi (hi : i ∈ s) (ht : disjoint (t₁ i) (t₂ i)) : + disjoint (s.pi t₁) (s.pi t₂) := +disjoint_left.2 $ λ h h₁ h₂, disjoint_left.1 ht (h₁ _ hi) (h₂ _ hi) + +section nonempty +variables [Π i, nonempty (α i)] + +lemma pi_eq_empty_iff' : s.pi t = ∅ ↔ ∃ i ∈ s, t i = ∅ := by simp [pi_eq_empty_iff] + +@[simp] lemma disjoint_pi : disjoint (s.pi t₁) (s.pi t₂) ↔ ∃ i ∈ s, disjoint (t₁ i) (t₂ i) := +by simp only [disjoint_iff_inter_eq_empty, ← pi_inter_distrib, pi_eq_empty_iff'] + +end nonempty + @[simp] lemma range_dcomp (f : Π i, α i → β i) : range (λ (g : Π i, α i), (λ i, f i (g i))) = pi univ (λ i, range (f i)) := begin diff --git a/src/order/sup_indep.lean b/src/order/sup_indep.lean index a91dab0d8ad13..572799ddcd38e 100644 --- a/src/order/sup_indep.lean +++ b/src/order/sup_indep.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Aaron Anderson, Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Aaron Anderson, Kevin Buzzard, Yaël Dillies, Eric Wieser -/ +import data.finset.sigma import data.finset.pairwise import data.finset.powerset import data.fintype.basic @@ -77,12 +78,45 @@ end lemma sup_indep.pairwise_disjoint (hs : s.sup_indep f) : (s : set ι).pairwise_disjoint f := λ a ha b hb hab, sup_singleton.subst $ hs (singleton_subset_iff.2 hb) ha $ not_mem_singleton.2 hab +lemma sup_indep.le_sup_iff (hs : s.sup_indep f) (hts : t ⊆ s) (hi : i ∈ s) (hf : ∀ i, f i ≠ ⊥) : + f i ≤ t.sup f ↔ i ∈ t := +begin + refine ⟨λ h, _, le_sup⟩, + by_contra hit, + exact hf i (disjoint_self.1 $ (hs hts hi hit).mono_right h), +end + /-- The RHS looks like the definition of `complete_lattice.independent`. -/ lemma sup_indep_iff_disjoint_erase [decidable_eq ι] : s.sup_indep f ↔ ∀ i ∈ s, disjoint (f i) ((s.erase i).sup f) := ⟨λ hs i hi, hs (erase_subset _ _) hi (not_mem_erase _ _), λ hs t ht i hi hit, (hs i hi).mono_right (sup_mono $ λ j hj, mem_erase.2 ⟨ne_of_mem_of_not_mem hj hit, ht hj⟩)⟩ +lemma sup_indep.image [decidable_eq ι] {s : finset ι'} {g : ι' → ι} (hs : s.sup_indep (f ∘ g)) : + (s.image g).sup_indep f := +begin + intros t ht i hi hit, + rw mem_image at hi, + obtain ⟨i, hi, rfl⟩ := hi, + haveI : decidable_eq ι' := classical.dec_eq _, + suffices hts : t ⊆ (s.erase i).image g, + { refine (sup_indep_iff_disjoint_erase.1 hs i hi).mono_right ((sup_mono hts).trans _), + rw sup_image }, + rintro j hjt, + obtain ⟨j, hj, rfl⟩ := mem_image.1 (ht hjt), + exact mem_image_of_mem _ (mem_erase.2 ⟨ne_of_apply_ne g (ne_of_mem_of_not_mem hjt hit), hj⟩), +end + +lemma sup_indep_map {s : finset ι'} {g : ι' ↪ ι} : (s.map g).sup_indep f ↔ s.sup_indep (f ∘ g) := +begin + refine ⟨λ hs t ht i hi hit, _, λ hs, _⟩, + { rw ←sup_map, + exact hs (map_subset_map.2 ht) ((mem_map' _).2 hi) (by rwa mem_map') }, + { classical, + rw map_eq_image, + exact hs.image } +end + @[simp] lemma sup_indep_pair [decidable_eq ι] {i j : ι} (hij : i ≠ j) : ({i, j} : finset ι).sup_indep f ↔ disjoint (f i) (f j) := ⟨λ h, h.pairwise_disjoint (by simp) (by simp) hij, λ h, begin @@ -117,7 +151,7 @@ begin exact sup_indep_pair this, end -lemma sup_indep.attach (hs : s.sup_indep f) : s.attach.sup_indep (f ∘ subtype.val) := +lemma sup_indep.attach (hs : s.sup_indep f) : s.attach.sup_indep (λ a, f a) := begin intros t ht i _ hi, classical, @@ -128,6 +162,18 @@ begin rwa subtype.ext hji at hj, end +@[simp] lemma sup_indep_attach : s.attach.sup_indep (λ a, f a) ↔ s.sup_indep f := +begin + refine ⟨λ h t ht i his hit, _, sup_indep.attach⟩, + classical, + convert h (filter_subset (λ i, (i : ι) ∈ t) _) (mem_attach _ ⟨i, ‹_›⟩) + (λ hi, hit $ by simpa using hi) using 1, + refine eq_of_forall_ge_iff _, + simp only [finset.sup_le_iff, mem_filter, mem_attach, true_and, function.comp_app, subtype.forall, + subtype.coe_mk], + exact λ a, forall_congr (λ j, ⟨λ h _, h, λ h hj, h (ht hj) hj⟩), +end + end lattice section distrib_lattice @@ -156,6 +202,58 @@ lemma sup_indep.bUnion [decidable_eq ι] {s : finset ι'} {g : ι' → finset ι (s.bUnion g).sup_indep f := by { rw ←sup_eq_bUnion, exact hs.sup hg } +/-- Bind operation for `sup_indep`. -/ +lemma sup_indep.sigma {β : ι → Type*} {s : finset ι} {g : Π i, finset (β i)} {f : sigma β → α} + (hs : s.sup_indep $ λ i, (g i).sup $ λ b, f ⟨i, b⟩) + (hg : ∀ i ∈ s, (g i).sup_indep $ λ b, f ⟨i, b⟩) : + (s.sigma g).sup_indep f := +begin + rintro t ht ⟨i, b⟩ hi hit, + rw finset.disjoint_sup_right, + rintro ⟨j, c⟩ hj, + have hbc := (ne_of_mem_of_not_mem hj hit).symm, + replace hj := ht hj, + rw mem_sigma at hi hj, + obtain rfl | hij := eq_or_ne i j, + { exact (hg _ hj.1).pairwise_disjoint hi.2 hj.2 (sigma_mk_injective.ne_iff.1 hbc) }, + { refine (hs.pairwise_disjoint hi.1 hj.1 hij).mono _ _, + { convert le_sup hi.2 }, + { convert le_sup hj.2 } } +end + +lemma sup_indep.product {s : finset ι} {t : finset ι'} {f : ι × ι' → α} + (hs : s.sup_indep $ λ i, t.sup $ λ i', f (i, i')) + (ht : t.sup_indep $ λ i', s.sup $ λ i, f (i, i')) : + (s.product t).sup_indep f := +begin + rintro u hu ⟨i, i'⟩ hi hiu, + rw finset.disjoint_sup_right, + rintro ⟨j, j'⟩ hj, + have hij := (ne_of_mem_of_not_mem hj hiu).symm, + replace hj := hu hj, + rw mem_product at hi hj, + obtain rfl | hij := eq_or_ne i j, + { refine (ht.pairwise_disjoint hi.2 hj.2 $ (prod.mk.inj_left _).ne_iff.1 hij).mono _ _, + { convert le_sup hi.1 }, + { convert le_sup hj.1 } }, + { refine (hs.pairwise_disjoint hi.1 hj.1 hij).mono _ _, + { convert le_sup hi.2 }, + { convert le_sup hj.2 } } +end + +lemma sup_indep_product_iff {s : finset ι} {t : finset ι'} {f : ι × ι' → α} : + (s.product t).sup_indep f ↔ + s.sup_indep (λ i, t.sup $ λ i', f (i, i')) ∧ t.sup_indep (λ i', s.sup $ λ i, f (i, i')) := +begin + refine ⟨_, λ h, h.1.product h.2⟩, + simp_rw sup_indep_iff_pairwise_disjoint, + refine (λ h, ⟨λ i hi j hj hij, _, λ i hi j hj hij, _⟩); + simp_rw [function.on_fun, finset.disjoint_sup_left, finset.disjoint_sup_right]; + intros i' hi' j' hj', + { exact h (mk_mem_product hi hi') (mk_mem_product hj hj') (ne_of_apply_ne prod.fst hij) }, + { exact h (mk_mem_product hi' hi) (mk_mem_product hj' hj) (ne_of_apply_ne prod.snd hij) } +end + end distrib_lattice end finset From 22c4d2ff43714b6ff724b2745ccfdc0f236a4a76 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 16 Jun 2023 11:48:58 +0000 Subject: [PATCH 1197/1291] feat(order/bounded_order): The lattice of complemented elements (#16267) Define `complementeds`, the subtype of complemented elements, and show that it is a complemented bounded distributive lattice. --- src/order/disjoint.lean | 83 ++++++++++++++++++++++++++++++++++++++++- 1 file changed, 82 insertions(+), 1 deletion(-) diff --git a/src/order/disjoint.lean b/src/order/disjoint.lean index bdff98cb1479d..85f97b552ba44 100644 --- a/src/order/disjoint.lean +++ b/src/order/disjoint.lean @@ -23,6 +23,8 @@ This file defines `disjoint`, `codisjoint`, and the `is_compl` predicate. -/ +open function + variable {α : Type*} section disjoint @@ -412,10 +414,32 @@ lemma eq_bot_of_top_is_compl (h : is_compl ⊤ x) : x = ⊥ := eq_top_of_bot_is_ end +section is_complemented +section lattice +variables [lattice α] [bounded_order α] + +/-- An element is *complemented* if it has a complement. -/ +def is_complemented (a : α) : Prop := ∃ b, is_compl a b + +lemma is_complemented_bot : is_complemented (⊥ : α) := ⟨⊤, is_compl_bot_top⟩ +lemma is_complemented_top : is_complemented (⊤ : α) := ⟨⊥, is_compl_top_bot⟩ + +end lattice + +variables [distrib_lattice α] [bounded_order α] {a b : α} + +lemma is_complemented.sup : is_complemented a → is_complemented b → is_complemented (a ⊔ b) := +λ ⟨a', ha⟩ ⟨b', hb⟩, ⟨a' ⊓ b', ha.sup_inf hb⟩ + +lemma is_complemented.inf : is_complemented a → is_complemented b → is_complemented (a ⊓ b) := +λ ⟨a', ha⟩ ⟨b', hb⟩, ⟨a' ⊔ b', ha.inf_sup hb⟩ + +end is_complemented + /-- A complemented bounded lattice is one where every element has a (not necessarily unique) complement. -/ class complemented_lattice (α) [lattice α] [bounded_order α] : Prop := -(exists_is_compl : ∀ (a : α), ∃ (b : α), is_compl a b) +(exists_is_compl : ∀ a : α, is_complemented a) export complemented_lattice (exists_is_compl) @@ -427,4 +451,61 @@ instance : complemented_lattice αᵒᵈ := end complemented_lattice +-- TODO: Define as a sublattice? +/-- The sublattice of complemented elements. -/ +@[reducible, derive partial_order] +def complementeds (α : Type*) [lattice α] [bounded_order α] : Type* := {a : α // is_complemented a} + +namespace complementeds +section lattice +variables [lattice α] [bounded_order α] {a b : complementeds α} + +instance has_coe_t : has_coe_t (complementeds α) α := ⟨subtype.val⟩ + +lemma coe_injective : injective (coe : complementeds α → α) := subtype.coe_injective + +@[simp, norm_cast] lemma coe_inj : (a : α) = b ↔ a = b := subtype.coe_inj +@[simp, norm_cast] lemma coe_le_coe : (a : α) ≤ b ↔ a ≤ b := by simp +@[simp, norm_cast] lemma coe_lt_coe : (a : α) < b ↔ a < b := iff.rfl + +instance : bounded_order (complementeds α) := +subtype.bounded_order is_complemented_bot is_complemented_top + +@[simp, norm_cast] lemma coe_bot : ((⊥ : complementeds α) : α) = ⊥ := rfl +@[simp, norm_cast] lemma coe_top : ((⊤ : complementeds α) : α) = ⊤ := rfl +@[simp] lemma mk_bot : (⟨⊥, is_complemented_bot⟩ : complementeds α) = ⊥ := rfl +@[simp] lemma mk_top : (⟨⊤, is_complemented_top⟩ : complementeds α) = ⊤ := rfl + +instance : inhabited (complementeds α) := ⟨⊥⟩ + +end lattice + +variables [distrib_lattice α] [bounded_order α] {a b : complementeds α} + +instance : has_sup (complementeds α) := ⟨λ a b, ⟨a ⊔ b, a.2.sup b.2⟩⟩ +instance : has_inf (complementeds α) := ⟨λ a b, ⟨a ⊓ b, a.2.inf b.2⟩⟩ + +@[simp, norm_cast] lemma coe_sup (a b : complementeds α) : (↑(a ⊔ b) : α) = a ⊔ b := rfl +@[simp, norm_cast] lemma coe_inf (a b : complementeds α) : (↑(a ⊓ b) : α) = a ⊓ b := rfl +@[simp] lemma mk_sup_mk {a b : α} (ha : is_complemented a) (hb : is_complemented b) : + (⟨a, ha⟩ ⊔ ⟨b, hb⟩ : complementeds α) = ⟨a ⊔ b, ha.sup hb⟩ := rfl +@[simp] lemma mk_inf_mk {a b : α} (ha : is_complemented a) (hb : is_complemented b) : + (⟨a, ha⟩ ⊓ ⟨b, hb⟩ : complementeds α) = ⟨a ⊓ b, ha.inf hb⟩ := rfl + +instance : distrib_lattice (complementeds α) := +complementeds.coe_injective.distrib_lattice _ coe_sup coe_inf + +@[simp, norm_cast] lemma disjoint_coe : disjoint (a : α) b ↔ disjoint a b := +by rw [disjoint_iff, disjoint_iff, ←coe_inf, ←coe_bot, coe_inj] + +@[simp, norm_cast] lemma codisjoint_coe : codisjoint (a : α) b ↔ codisjoint a b := +by rw [codisjoint_iff, codisjoint_iff, ←coe_sup, ←coe_top, coe_inj] + +@[simp, norm_cast] lemma is_compl_coe : is_compl (a : α) b ↔ is_compl a b := +by simp_rw [is_compl_iff, disjoint_coe, codisjoint_coe] + +instance : complemented_lattice (complementeds α) := +⟨λ ⟨a, b, h⟩, ⟨⟨b, a, h.symm⟩, is_compl_coe.1 h⟩⟩ + +end complementeds end is_compl From 893964fc28cefbcffc7cb784ed00a2895b4e65cf Mon Sep 17 00:00:00 2001 From: jakelev Date: Fri, 16 Jun 2023 17:36:19 +0000 Subject: [PATCH 1198/1291] feat(algebra/homology/local_cohomology): Add equivalence of local cohomology along cofinal diagrams (#19105) Add the natural equivalence of local cohomology along cofinal diagrams of ideals. Co-authored-by: Scott Morrison --- src/algebra/homology/local_cohomology.lean | 118 ++++++++++++++++----- 1 file changed, 94 insertions(+), 24 deletions(-) diff --git a/src/algebra/homology/local_cohomology.lean b/src/algebra/homology/local_cohomology.lean index ea2ccd6ee952a..2812460b2b143 100644 --- a/src/algebra/homology/local_cohomology.lean +++ b/src/algebra/homology/local_cohomology.lean @@ -7,7 +7,8 @@ import ring_theory.ideal.basic import algebra.category.Module.colimits import algebra.category.Module.projective import category_theory.abelian.ext -import ring_theory.finiteness +import category_theory.limits.final +import ring_theory.noetherian /-! # Local cohomology. @@ -42,7 +43,6 @@ local cohomology, local cohomology modules * the right-derived functor definition * the characterization as the limit of Koszul homology * the characterization as the cohomology of a Cech-like complex -* Prove that local cohomology depends only on the radical of the ideal * Establish long exact sequence(s) in local cohomology -/ @@ -52,8 +52,7 @@ open category_theory.limits noncomputable theory -universes u v - +universes u v v' namespace local_cohomology @@ -90,10 +89,9 @@ of `R`, represented as a functor `I`. -/ In this definition we do not assume any special property of the diagram `I`, but the relevant case will be where `I` is (cofinal with) the diagram of powers of a single given ideal. -Below, we give two equivalent (to be shown) definitions of the usual local cohomology with support +Below, we give two equivalent definitions of the usual local cohomology with support in an ideal `J`, `local_cohomology` and `local_cohomology.of_self_le_radical`. -TODO: Show that any functor cofinal with `I` gives the same result. -/ def of_diagram (I : D ⥤ ideal R) (i : ℕ) : Module.{max u v} R ⥤ Module.{max u v} R := @@ -101,6 +99,23 @@ colimit (diagram.{(max u v) v} I i) end +section +variables {R : Type max u v v'} [comm_ring R] {D : Type v} [small_category D] + +variables {E : Type v'} [small_category E] + (I' : E ⥤ D) (I : D ⥤ ideal R) + +/-- Local cohomology along a composition of diagrams. -/ +def diagram_comp (i : ℕ) : diagram (I' ⋙ I) i ≅ I'.op ⋙ (diagram I i) := iso.refl _ + +/-- Local cohomology agrees along precomposition with a cofinal diagram. -/ +def iso_of_final [functor.initial I'] (i : ℕ) : + of_diagram.{(max u v) v'} (I' ⋙ I) i ≅ of_diagram.{(max u v') v} I i := +(has_colimit.iso_of_nat_iso (diagram_comp _ _ _)) +≪≫ (functor.final.colimit_iso _ _) + +end + section diagrams variables {R : Type u} [comm_ring R] @@ -139,28 +154,31 @@ variables {R : Type u} [comm_ring R] a commutative ring `R` with support in the ideal `J` of `R`, defined as the direct limit of `Ext^i(R/J^t, M)` over all powers `t : ℕ`. -/ def local_cohomology (J : ideal R) (i : ℕ) : Module.{u} R ⥤ Module.{u} R := - of_diagram (ideal_powers_diagram J) i +of_diagram (ideal_powers_diagram J) i /-- Local cohomology as the direct limit of `Ext^i(R/J', M)` over *all* ideals `J'` with radical containing `J`. -/ -def local_cohomology.of_self_le_radical (J : ideal R) (i : ℕ) : - Module.{u} R ⥤ Module.{u} R := +def local_cohomology.of_self_le_radical (J : ideal R) (i : ℕ) : Module.{u} R ⥤ Module.{u} R := of_diagram.{u} (self_le_radical_diagram.{u} J) i -/- TODO: Construct `local_cohomology J i ≅ local_cohomology.of_self_le_radical J i`. Use this to -show that local cohomology depends only on `J.radical`. -/ end models_for_local_cohomology +namespace local_cohomology +/-! +Showing equivalence of different definitions of local cohomology. + * `local_cohomology.iso_self_le_radical` gives the isomorphism + `local_cohomology J i ≅ local_cohomology.of_self_le_radical J i` + * `local_cohomology.iso_of_same_radical` gives the isomorphism + `local_cohomology J i ≅ local_cohomology K i` when `J.radical = K.radical`. +-/ section local_cohomology_equiv -open local_cohomology - -variables {R : Type u} [comm_ring R] (I J : ideal R) +variables {R : Type u} [comm_ring R] /-- Lifting `ideal_powers_diagram J` from a diagram valued in `ideals R` to a diagram valued in `self_le_radical J`. -/ -def local_cohomology.ideal_powers_to_self_le_radical (J : ideal R) : ℕᵒᵖ ⥤ self_le_radical J := +def ideal_powers_to_self_le_radical (J : ideal R) : ℕᵒᵖ ⥤ self_le_radical J := full_subcategory.lift _ (ideal_powers_diagram J) (λ k, begin change _ ≤ (J^(unop k)).radical, @@ -169,17 +187,12 @@ full_subcategory.lift _ (ideal_powers_diagram J) { simp only [J.radical_pow _ n.succ_pos, ideal.le_radical] }, end) -/-- The composition with the inclusion into `ideals R` is isomorphic to `ideal_powers_diagram J`. -/ -def local_cohomology.ideal_powers_to_self_le_radical_comp_inclusion (J : ideal R) : -local_cohomology.ideal_powers_to_self_le_radical J ⋙ self_le_radical_diagram J - ≅ ideal_powers_diagram J := - full_subcategory.lift_comp_inclusion _ _ _ - -/-- The lemma below essentially says that `ideal_powers_to_self_le_radical I` is initial in -`self_le_radical_diagram I`. +variables {I J K : ideal R} +/-- PORTING NOTE: This lemma should probably be moved to `ring_theory/finiteness.lean` -to be near `ideal.exists_radical_pow_le_of_fg`, which it generalizes. -/ +to be near `ideal.exists_radical_pow_le_of_fg`, which it generalizes. +-/ lemma ideal.exists_pow_le_of_le_radical_of_fg (hIJ : I ≤ J.radical) (hJ : J.radical.fg) : ∃ (k : ℕ), I^k ≤ J := begin @@ -189,4 +202,61 @@ begin ... ≤ J : hk, end +/-- The diagram of powers of `J` is initial in the diagram of all ideals with +radical containing `J`. This uses noetherianness. -/ +instance ideal_powers_initial [hR : is_noetherian R R] : + functor.initial (ideal_powers_to_self_le_radical J) := +{ out := λ J', begin + apply @zigzag_is_connected _ _ _, + { intros j1 j2, + apply relation.refl_trans_gen.single, + -- The inclusions `J^n1 ≤ J'` and `J^n2 ≤ J'` always form a triangle, based on + -- which exponent is larger. + cases le_total (unop j1.left) (unop j2.left) with h, + right, exact ⟨costructured_arrow.hom_mk (hom_of_le h).op (of_as_true trivial)⟩, + left, exact ⟨costructured_arrow.hom_mk (hom_of_le h).op (of_as_true trivial)⟩ }, + { obtain ⟨k, hk⟩ := ideal.exists_pow_le_of_le_radical_of_fg J'.2 + (is_noetherian_def.mp hR _), + exact ⟨costructured_arrow.mk (⟨⟨hk⟩⟩ : (ideal_powers_to_self_le_radical J).obj (op k) ⟶ J')⟩}, + end } + +/-- Local cohomology (defined in terms of powers of `J`) agrees with local +cohomology computed over all ideals with radical containing `J`. -/ +def iso_self_le_radical (J : ideal R) [is_noetherian R R] (i : ℕ) : + local_cohomology.of_self_le_radical J i ≅ local_cohomology J i := +(local_cohomology.iso_of_final.{u u 0} + (ideal_powers_to_self_le_radical J) (self_le_radical_diagram J) i).symm +≪≫ has_colimit.iso_of_nat_iso (iso.refl _) + +/-- Casting from the full subcategory of ideals with radical containing `J` to the full +subcategory of ideals with radical containing `K`. -/ +def self_le_radical.cast (hJK : J.radical = K.radical) : + self_le_radical J ⥤ self_le_radical K := +full_subcategory.map (λ L hL, begin + rw ← ideal.radical_le_radical_iff at ⊢ hL, + exact hJK.symm.trans_le hL, + end) + +-- TODO generalize this to the equivalence of full categories for any `iff`. +instance self_le_radical.cast_is_equivalence (hJK : J.radical = K.radical) : + is_equivalence (self_le_radical.cast hJK) := +{ inverse := self_le_radical.cast hJK.symm, + unit_iso := by tidy, + counit_iso := by tidy } + +/-- The natural isomorphism between local cohomology defined using the `of_self_le_radical` +diagram, assuming `J.radical = K.radical`. -/ +def self_le_radical.iso_of_same_radical (hJK : J.radical = K.radical) (i : ℕ) : + of_self_le_radical J i ≅ of_self_le_radical K i := +(iso_of_final.{u u u} (self_le_radical.cast hJK.symm) _ _).symm + +/-- Local cohomology agrees on ideals with the same radical. -/ +def iso_of_same_radical [is_noetherian R R] (hJK : J.radical = K.radical) (i : ℕ) : + local_cohomology J i ≅ local_cohomology K i := +(iso_self_le_radical J i).symm +≪≫ self_le_radical.iso_of_same_radical hJK i +≪≫ iso_self_le_radical K i + end local_cohomology_equiv + +end local_cohomology From 2a0ce625dbb0ffbc7d1316597de0b25c1ec75303 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Sat, 17 Jun 2023 06:21:17 +0000 Subject: [PATCH 1199/1291] chore(*): add mathlib4 synchronization comments (#19196) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.category.BoolRing` * `algebra.continued_fractions.computation.approximation_corollaries` * `category_theory.monoidal.internal.functor_category` * `field_theory.galois` * `field_theory.minpoly.is_integrally_closed` * `field_theory.primitive_element` * `linear_algebra.matrix.ldl` * `number_theory.diophantine_approximation` * `number_theory.pell` * `number_theory.primes_congruent_one` * `order.category.FinBoolAlg` * `probability.kernel.composition` * `probability.kernel.cond_cdf` * `probability.kernel.invariance` * `ring_theory.polynomial.cyclotomic.eval` * `ring_theory.polynomial.cyclotomic.expand` * `ring_theory.polynomial.cyclotomic.roots` * `ring_theory.polynomial.gauss_lemma` * `ring_theory.polynomial.selmer` * `ring_theory.ring_hom.finite_type` * `ring_theory.roots_of_unity.minpoly` * `topology.category.UniformSpace` * `topology.continuous_function.zero_at_infty` * `topology.sheaves.operations` --- src/algebra/category/BoolRing.lean | 3 +++ .../computation/approximation_corollaries.lean | 3 +++ src/category_theory/monoidal/internal/functor_category.lean | 3 +++ src/field_theory/galois.lean | 3 +++ src/field_theory/minpoly/is_integrally_closed.lean | 3 +++ src/field_theory/primitive_element.lean | 3 +++ src/linear_algebra/matrix/ldl.lean | 3 +++ src/number_theory/diophantine_approximation.lean | 3 +++ src/number_theory/pell.lean | 3 +++ src/number_theory/primes_congruent_one.lean | 3 +++ src/order/category/FinBoolAlg.lean | 3 +++ src/probability/kernel/composition.lean | 3 +++ src/probability/kernel/cond_cdf.lean | 3 +++ src/probability/kernel/invariance.lean | 3 +++ src/ring_theory/polynomial/cyclotomic/eval.lean | 3 +++ src/ring_theory/polynomial/cyclotomic/expand.lean | 3 +++ src/ring_theory/polynomial/cyclotomic/roots.lean | 3 +++ src/ring_theory/polynomial/gauss_lemma.lean | 3 +++ src/ring_theory/polynomial/selmer.lean | 3 +++ src/ring_theory/ring_hom/finite_type.lean | 3 +++ src/ring_theory/roots_of_unity/minpoly.lean | 3 +++ src/topology/category/UniformSpace.lean | 3 +++ src/topology/continuous_function/zero_at_infty.lean | 3 +++ src/topology/sheaves/operations.lean | 3 +++ 24 files changed, 72 insertions(+) diff --git a/src/algebra/category/BoolRing.lean b/src/algebra/category/BoolRing.lean index 70ab2bdec7803..9ec49984dc885 100644 --- a/src/algebra/category/BoolRing.lean +++ b/src/algebra/category/BoolRing.lean @@ -10,6 +10,9 @@ import order.category.BoolAlg /-! # The category of Boolean rings +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines `BoolRing`, the category of Boolean rings. ## TODO diff --git a/src/algebra/continued_fractions/computation/approximation_corollaries.lean b/src/algebra/continued_fractions/computation/approximation_corollaries.lean index df71b81f9db55..0e8d5fab9f97a 100644 --- a/src/algebra/continued_fractions/computation/approximation_corollaries.lean +++ b/src/algebra/continued_fractions/computation/approximation_corollaries.lean @@ -12,6 +12,9 @@ import topology.order.basic /-! # Corollaries From Approximation Lemmas (`algebra.continued_fractions.computation.approximations`) +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Summary We show that the generalized_continued_fraction given by `generalized_continued_fraction.of` in fact diff --git a/src/category_theory/monoidal/internal/functor_category.lean b/src/category_theory/monoidal/internal/functor_category.lean index 11f40f0aa9533..8e39949a504ef 100644 --- a/src/category_theory/monoidal/internal/functor_category.lean +++ b/src/category_theory/monoidal/internal/functor_category.lean @@ -9,6 +9,9 @@ import category_theory.monoidal.functor_category /-! # `Mon_ (C ⥤ D) ≌ C ⥤ Mon_ D` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + When `D` is a monoidal category, monoid objects in `C ⥤ D` are the same thing as functors from `C` into the monoid objects of `D`. diff --git a/src/field_theory/galois.lean b/src/field_theory/galois.lean index 1efa971954436..f1bc6f0db9643 100644 --- a/src/field_theory/galois.lean +++ b/src/field_theory/galois.lean @@ -11,6 +11,9 @@ import group_theory.group_action.fixing_subgroup /-! # Galois Extensions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define Galois extensions as extensions which are both separable and normal. ## Main definitions diff --git a/src/field_theory/minpoly/is_integrally_closed.lean b/src/field_theory/minpoly/is_integrally_closed.lean index 837cc2ad4face..6a10a09a9738e 100644 --- a/src/field_theory/minpoly/is_integrally_closed.lean +++ b/src/field_theory/minpoly/is_integrally_closed.lean @@ -10,6 +10,9 @@ import ring_theory.polynomial.gauss_lemma /-! # Minimal polynomials over a GCD monoid +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file specializes the theory of minpoly to the case of an algebra over a GCD monoid. ## Main results diff --git a/src/field_theory/primitive_element.lean b/src/field_theory/primitive_element.lean index 0b7587117823a..cf57da928a8f7 100644 --- a/src/field_theory/primitive_element.lean +++ b/src/field_theory/primitive_element.lean @@ -12,6 +12,9 @@ import ring_theory.integral_domain /-! # Primitive Element Theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove the primitive element theorem. ## Main results diff --git a/src/linear_algebra/matrix/ldl.lean b/src/linear_algebra/matrix/ldl.lean index f346a739ba6b3..55458ac40f91d 100644 --- a/src/linear_algebra/matrix/ldl.lean +++ b/src/linear_algebra/matrix/ldl.lean @@ -8,6 +8,9 @@ import linear_algebra.matrix.pos_def /-! # LDL decomposition +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves the LDL-decomposition of matricies: Any positive definite matrix `S` can be decomposed as `S = LDLᴴ` where `L` is a lower-triangular matrix and `D` is a diagonal matrix. diff --git a/src/number_theory/diophantine_approximation.lean b/src/number_theory/diophantine_approximation.lean index 556dc53c8297a..772f7e19881e2 100644 --- a/src/number_theory/diophantine_approximation.lean +++ b/src/number_theory/diophantine_approximation.lean @@ -14,6 +14,9 @@ import tactic.basic /-! # Diophantine Approximation +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The first part of this file gives proofs of various versions of **Dirichlet's approximation theorem** and its important consequence that when $\xi$ is an irrational real number, then there are infinitely many rationals $x/y$ (in lowest terms) diff --git a/src/number_theory/pell.lean b/src/number_theory/pell.lean index 53fde0c78ae84..a9f498d0bd0fd 100644 --- a/src/number_theory/pell.lean +++ b/src/number_theory/pell.lean @@ -12,6 +12,9 @@ import number_theory.zsqrtd.basic /-! # Pell's Equation +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + *Pell's Equation* is the equation $x^2 - d y^2 = 1$, where $d$ is a positive integer that is not a square, and one is interested in solutions in integers $x$ and $y$. diff --git a/src/number_theory/primes_congruent_one.lean b/src/number_theory/primes_congruent_one.lean index 21e0107bb3e66..68d6d0d28d4d5 100644 --- a/src/number_theory/primes_congruent_one.lean +++ b/src/number_theory/primes_congruent_one.lean @@ -9,6 +9,9 @@ import ring_theory.polynomial.cyclotomic.eval /-! # Primes congruent to one +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We prove that, for any positive `k : ℕ`, there are infinitely many primes `p` such that `p ≡ 1 [MOD k]`. -/ diff --git a/src/order/category/FinBoolAlg.lean b/src/order/category/FinBoolAlg.lean index 5082fe335a6ed..ad889807c39f8 100644 --- a/src/order/category/FinBoolAlg.lean +++ b/src/order/category/FinBoolAlg.lean @@ -11,6 +11,9 @@ import order.hom.complete_lattice /-! # The category of finite boolean algebras +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines `FinBoolAlg`, the category of finite boolean algebras. ## TODO diff --git a/src/probability/kernel/composition.lean b/src/probability/kernel/composition.lean index 17cc45818f356..4142e94a07ce4 100644 --- a/src/probability/kernel/composition.lean +++ b/src/probability/kernel/composition.lean @@ -9,6 +9,9 @@ import probability.kernel.measurable_integral /-! # Product and composition of kernels +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define * the composition-product `κ ⊗ₖ η` of two s-finite kernels `κ : kernel α β` and `η : kernel (α × β) γ`, a kernel from `α` to `β × γ`. diff --git a/src/probability/kernel/cond_cdf.lean b/src/probability/kernel/cond_cdf.lean index e9478de11a3c7..6a8170f2918d8 100644 --- a/src/probability/kernel/cond_cdf.lean +++ b/src/probability/kernel/cond_cdf.lean @@ -10,6 +10,9 @@ import measure_theory.decomposition.radon_nikodym /-! # Conditional cumulative distribution function +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given `ρ : measure (α × ℝ)`, we define the conditional cumulative distribution function (conditional cdf) of `ρ`. It is a function `cond_cdf ρ : α → ℝ → ℝ` such that if `ρ` is a finite measure, then for all `a : α` `cond_cdf ρ a` is monotone and right-continuous with limit 0 at -∞ diff --git a/src/probability/kernel/invariance.lean b/src/probability/kernel/invariance.lean index a33a340bb8e74..478deaee8768a 100644 --- a/src/probability/kernel/invariance.lean +++ b/src/probability/kernel/invariance.lean @@ -8,6 +8,9 @@ import probability.kernel.composition /-! # Invariance of measures along a kernel +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We say that a measure `μ` is invariant with respect to a kernel `κ` if its push-forward along the kernel `μ.bind κ` is the same measure. diff --git a/src/ring_theory/polynomial/cyclotomic/eval.lean b/src/ring_theory/polynomial/cyclotomic/eval.lean index 84bb6f3300a52..266c5c6c3698c 100644 --- a/src/ring_theory/polynomial/cyclotomic/eval.lean +++ b/src/ring_theory/polynomial/cyclotomic/eval.lean @@ -12,6 +12,9 @@ import analysis.complex.arg /-! # Evaluating cyclotomic polynomials + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. This file states some results about evaluating cyclotomic polynomials in various different ways. ## Main definitions * `polynomial.eval(₂)_one_cyclotomic_prime(_pow)`: `eval 1 (cyclotomic p^k R) = p`. diff --git a/src/ring_theory/polynomial/cyclotomic/expand.lean b/src/ring_theory/polynomial/cyclotomic/expand.lean index 83ae1209b818f..f79fa27ec789a 100644 --- a/src/ring_theory/polynomial/cyclotomic/expand.lean +++ b/src/ring_theory/polynomial/cyclotomic/expand.lean @@ -9,6 +9,9 @@ import ring_theory.polynomial.cyclotomic.roots /-! # Cyclotomic polynomials and `expand`. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We gather results relating cyclotomic polynomials and `expand`. ## Main results diff --git a/src/ring_theory/polynomial/cyclotomic/roots.lean b/src/ring_theory/polynomial/cyclotomic/roots.lean index 8fb945ebf74d6..41e73830f5ffc 100644 --- a/src/ring_theory/polynomial/cyclotomic/roots.lean +++ b/src/ring_theory/polynomial/cyclotomic/roots.lean @@ -10,6 +10,9 @@ import ring_theory.roots_of_unity.minpoly /-! # Roots of cyclotomic polynomials. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We gather results about roots of cyclotomic polynomials. In particular we show in `polynomial.cyclotomic_eq_minpoly` that `cyclotomic n R` is the minimal polynomial of a primitive root of unity. diff --git a/src/ring_theory/polynomial/gauss_lemma.lean b/src/ring_theory/polynomial/gauss_lemma.lean index ec630a2b49f90..ee3651e4a3d0e 100644 --- a/src/ring_theory/polynomial/gauss_lemma.lean +++ b/src/ring_theory/polynomial/gauss_lemma.lean @@ -12,6 +12,9 @@ import ring_theory.integrally_closed /-! # Gauss's Lemma +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Gauss's Lemma is one of a few results pertaining to irreducibility of primitive polynomials. ## Main Results diff --git a/src/ring_theory/polynomial/selmer.lean b/src/ring_theory/polynomial/selmer.lean index 4d449cb563af7..ba8ec36036055 100644 --- a/src/ring_theory/polynomial/selmer.lean +++ b/src/ring_theory/polynomial/selmer.lean @@ -11,6 +11,9 @@ import tactic.linear_combination /-! # Irreducibility of Selmer Polynomials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves irreducibility of the Selmer polynomials `X ^ n - X - 1`. ## Main results diff --git a/src/ring_theory/ring_hom/finite_type.lean b/src/ring_theory/ring_hom/finite_type.lean index 6f6f5ab0c978a..3d5d2f9b43b15 100644 --- a/src/ring_theory/ring_hom/finite_type.lean +++ b/src/ring_theory/ring_hom/finite_type.lean @@ -10,6 +10,9 @@ import ring_theory.localization.inv_submonoid # The meta properties of finite-type ring homomorphisms. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The main result is `ring_hom.finite_is_local`. -/ diff --git a/src/ring_theory/roots_of_unity/minpoly.lean b/src/ring_theory/roots_of_unity/minpoly.lean index 5ab77bce66d05..4c5cf407e5a63 100644 --- a/src/ring_theory/roots_of_unity/minpoly.lean +++ b/src/ring_theory/roots_of_unity/minpoly.lean @@ -10,6 +10,9 @@ import field_theory.minpoly.is_integrally_closed /-! # Minimal polynomial of roots of unity +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We gather several results about minimal polynomial of root of unity. ## Main results diff --git a/src/topology/category/UniformSpace.lean b/src/topology/category/UniformSpace.lean index c77bcd9e646dd..226cd56608165 100644 --- a/src/topology/category/UniformSpace.lean +++ b/src/topology/category/UniformSpace.lean @@ -12,6 +12,9 @@ import topology.uniform_space.completion /-! # The category of uniform spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We construct the category of uniform spaces, show that the complete separated uniform spaces form a reflective subcategory, and hence possess all limits that uniform spaces do. diff --git a/src/topology/continuous_function/zero_at_infty.lean b/src/topology/continuous_function/zero_at_infty.lean index f85b63491ed92..60e0bb52fd6df 100644 --- a/src/topology/continuous_function/zero_at_infty.lean +++ b/src/topology/continuous_function/zero_at_infty.lean @@ -9,6 +9,9 @@ import topology.continuous_function.cocompact_map /-! # Continuous functions vanishing at infinity +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The type of continuous functions vanishing at infinity. When the domain is compact `C(α, β) ≃ C₀(α, β)` via the identity map. When the codomain is a metric space, every continuous map which vanishes at infinity is a bounded continuous function. When the domain is a locally diff --git a/src/topology/sheaves/operations.lean b/src/topology/sheaves/operations.lean index 8358d44973eb9..c4ce3bb1979d9 100644 --- a/src/topology/sheaves/operations.lean +++ b/src/topology/sheaves/operations.lean @@ -12,6 +12,9 @@ import topology.sheaves.stalks # Operations on sheaves +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definition - `submonoid_presheaf` : A subpresheaf with a submonoid structure on each of the components. From e160cefedc932ce41c7049bf0c4b0f061d06216e Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Mon, 19 Jun 2023 06:39:32 +0000 Subject: [PATCH 1200/1291] chore(*): add mathlib4 synchronization comments (#19200) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.continued_fractions.computation.terminates_iff_rat` * `category_theory.enriched.basic` * `category_theory.monoidal.internal.limits` * `field_theory.krull_topology` * `measure_theory.function.conditional_expectation.basic` * `measure_theory.function.conditional_expectation.condexp_L1` * `measure_theory.function.conditional_expectation.condexp_L2` * `measure_theory.function.conditional_expectation.indicator` * `measure_theory.function.conditional_expectation.real` * `number_theory.legendre_symbol.basic` * `number_theory.legendre_symbol.quadratic_char.basic` * `probability.conditional_expectation` * `probability.notation` * `probability.process.adapted` * `probability.process.filtration` * `probability.process.stopping` --- .../continued_fractions/computation/terminates_iff_rat.lean | 3 +++ src/category_theory/enriched/basic.lean | 3 +++ src/category_theory/monoidal/internal/limits.lean | 3 +++ src/field_theory/krull_topology.lean | 3 +++ src/measure_theory/function/conditional_expectation/basic.lean | 3 +++ .../function/conditional_expectation/condexp_L1.lean | 3 +++ .../function/conditional_expectation/condexp_L2.lean | 3 +++ .../function/conditional_expectation/indicator.lean | 3 +++ src/measure_theory/function/conditional_expectation/real.lean | 3 +++ src/number_theory/legendre_symbol/basic.lean | 3 +++ src/number_theory/legendre_symbol/quadratic_char/basic.lean | 3 +++ src/probability/conditional_expectation.lean | 3 +++ src/probability/notation.lean | 3 +++ src/probability/process/adapted.lean | 3 +++ src/probability/process/filtration.lean | 3 +++ src/probability/process/stopping.lean | 3 +++ 16 files changed, 48 insertions(+) diff --git a/src/algebra/continued_fractions/computation/terminates_iff_rat.lean b/src/algebra/continued_fractions/computation/terminates_iff_rat.lean index d4977bd208740..a559d3f765016 100644 --- a/src/algebra/continued_fractions/computation/terminates_iff_rat.lean +++ b/src/algebra/continued_fractions/computation/terminates_iff_rat.lean @@ -9,6 +9,9 @@ import data.rat.floor /-! # Termination of Continued Fraction Computations (`gcf.of`) +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Summary We show that the continued fraction for a value `v`, as defined in `algebra.continued_fractions.computation.basic`, terminates if and only if `v` corresponds to a diff --git a/src/category_theory/enriched/basic.lean b/src/category_theory/enriched/basic.lean index 1510eb07d5e5d..a7eea5c06bb38 100644 --- a/src/category_theory/enriched/basic.lean +++ b/src/category_theory/enriched/basic.lean @@ -11,6 +11,9 @@ import tactic.apply_fun /-! # Enriched categories +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We set up the basic theory of `V`-enriched categories, for `V` an arbitrary monoidal category. diff --git a/src/category_theory/monoidal/internal/limits.lean b/src/category_theory/monoidal/internal/limits.lean index 45f02704db445..03df17a27c861 100644 --- a/src/category_theory/monoidal/internal/limits.lean +++ b/src/category_theory/monoidal/internal/limits.lean @@ -10,6 +10,9 @@ import category_theory.limits.preserves.basic /-! # Limits of monoid objects. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + If `C` has limits, so does `Mon_ C`, and the forgetful functor preserves these limits. (This could potentially replace many individual constructions for concrete categories, diff --git a/src/field_theory/krull_topology.lean b/src/field_theory/krull_topology.lean index 0901d5e87b26a..03894aa273074 100644 --- a/src/field_theory/krull_topology.lean +++ b/src/field_theory/krull_topology.lean @@ -12,6 +12,9 @@ import tactic.by_contra /-! # Krull topology +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define the Krull topology on `L ≃ₐ[K] L` for an arbitrary field extension `L/K`. In order to do this, we first define a `group_filter_basis` on `L ≃ₐ[K] L`, whose sets are `E.fixing_subgroup` for all intermediate fields `E` with `E/K` finite dimensional. diff --git a/src/measure_theory/function/conditional_expectation/basic.lean b/src/measure_theory/function/conditional_expectation/basic.lean index 404015d07d653..dab82b38e673d 100644 --- a/src/measure_theory/function/conditional_expectation/basic.lean +++ b/src/measure_theory/function/conditional_expectation/basic.lean @@ -7,6 +7,9 @@ import measure_theory.function.conditional_expectation.condexp_L1 /-! # Conditional expectation +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We build the conditional expectation of an integrable function `f` with value in a Banach space with respect to a measure `μ` (defined on a measurable space structure `m0`) and a measurable space structure `m` with `hm : m ≤ m0` (a sub-sigma-algebra). This is an `m`-strongly measurable diff --git a/src/measure_theory/function/conditional_expectation/condexp_L1.lean b/src/measure_theory/function/conditional_expectation/condexp_L1.lean index 1368135fae32e..1f5729691234e 100644 --- a/src/measure_theory/function/conditional_expectation/condexp_L1.lean +++ b/src/measure_theory/function/conditional_expectation/condexp_L1.lean @@ -7,6 +7,9 @@ import measure_theory.function.conditional_expectation.condexp_L2 /-! # Conditional expectation in L1 +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains two more steps of the construction of the conditional expectation, which is completed in `measure_theory.function.conditional_expectation.basic`. See that file for a description of the full process. diff --git a/src/measure_theory/function/conditional_expectation/condexp_L2.lean b/src/measure_theory/function/conditional_expectation/condexp_L2.lean index fc1d507a730ff..b099139680fd1 100644 --- a/src/measure_theory/function/conditional_expectation/condexp_L2.lean +++ b/src/measure_theory/function/conditional_expectation/condexp_L2.lean @@ -8,6 +8,9 @@ import measure_theory.function.conditional_expectation.unique /-! # Conditional expectation in L2 +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains one step of the construction of the conditional expectation, which is completed in `measure_theory.function.conditional_expectation.basic`. See that file for a description of the full process. diff --git a/src/measure_theory/function/conditional_expectation/indicator.lean b/src/measure_theory/function/conditional_expectation/indicator.lean index 4925425dcf5c1..75267db6ea2d7 100644 --- a/src/measure_theory/function/conditional_expectation/indicator.lean +++ b/src/measure_theory/function/conditional_expectation/indicator.lean @@ -10,6 +10,9 @@ import measure_theory.function.conditional_expectation.basic # Conditional expectation of indicator functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves some results about the conditional expectation of an indicator function and as a corollary, also proves several results about the behaviour of the conditional expectation on a restricted measure. diff --git a/src/measure_theory/function/conditional_expectation/real.lean b/src/measure_theory/function/conditional_expectation/real.lean index 1d5a8929a2e18..ffdda9fc88a67 100644 --- a/src/measure_theory/function/conditional_expectation/real.lean +++ b/src/measure_theory/function/conditional_expectation/real.lean @@ -12,6 +12,9 @@ import measure_theory.decomposition.radon_nikodym # Conditional expectation of real-valued functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves some results regarding the conditional expectation of real-valued functions. ## Main results diff --git a/src/number_theory/legendre_symbol/basic.lean b/src/number_theory/legendre_symbol/basic.lean index 545e7e6149a45..4427fe06ad800 100644 --- a/src/number_theory/legendre_symbol/basic.lean +++ b/src/number_theory/legendre_symbol/basic.lean @@ -8,6 +8,9 @@ import number_theory.legendre_symbol.quadratic_char.basic /-! # Legendre symbol +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains results about Legendre symbols. We define the Legendre symbol $\Bigl(\frac{a}{p}\Bigr)$ as `legendre_sym p a`. diff --git a/src/number_theory/legendre_symbol/quadratic_char/basic.lean b/src/number_theory/legendre_symbol/quadratic_char/basic.lean index 0d87db3f03a63..8d64425b47fff 100644 --- a/src/number_theory/legendre_symbol/quadratic_char/basic.lean +++ b/src/number_theory/legendre_symbol/quadratic_char/basic.lean @@ -10,6 +10,9 @@ import field_theory.finite.basic /-! # Quadratic characters of finite fields +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the quadratic character on a finite field `F` and proves some basic statements about it. diff --git a/src/probability/conditional_expectation.lean b/src/probability/conditional_expectation.lean index d77627101bee3..beedf14825e8f 100644 --- a/src/probability/conditional_expectation.lean +++ b/src/probability/conditional_expectation.lean @@ -11,6 +11,9 @@ import measure_theory.function.conditional_expectation.basic # Probabilistic properties of the conditional expectation +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains some properties about the conditional expectation which does not belong in the main conditional expectation file. diff --git a/src/probability/notation.lean b/src/probability/notation.lean index 5e9d1e14b5f1c..5b6d45c1d4f42 100644 --- a/src/probability/notation.lean +++ b/src/probability/notation.lean @@ -8,6 +8,9 @@ import measure_theory.function.conditional_expectation.basic /-! # Notations for probability theory +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the following notations, for functions `X,Y`, measures `P, Q` defined on a measurable space `m0`, and another measurable space structure `m` with `hm : m ≤ m0`, - `P[X] = ∫ a, X a ∂P` diff --git a/src/probability/process/adapted.lean b/src/probability/process/adapted.lean index 3cb162bc3b916..a483a31db3c44 100644 --- a/src/probability/process/adapted.lean +++ b/src/probability/process/adapted.lean @@ -9,6 +9,9 @@ import topology.instances.discrete /-! # Adapted and progressively measurable processes +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines some standard definition from the theory of stochastic processes including filtrations and stopping times. These definitions are used to model the amount of information at a specific time and are the first step in formalizing stochastic processes. diff --git a/src/probability/process/filtration.lean b/src/probability/process/filtration.lean index 7a3c14f17260b..9bcf5a38d2a62 100644 --- a/src/probability/process/filtration.lean +++ b/src/probability/process/filtration.lean @@ -8,6 +8,9 @@ import measure_theory.function.conditional_expectation.real /-! # Filtrations +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines filtrations of a measurable space and σ-finite filtrations. ## Main definitions diff --git a/src/probability/process/stopping.lean b/src/probability/process/stopping.lean index be11247dd49c5..cec8a181ef003 100644 --- a/src/probability/process/stopping.lean +++ b/src/probability/process/stopping.lean @@ -9,6 +9,9 @@ import probability.process.adapted /-! # Stopping times, stopped processes and stopped values +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Definition and properties of stopping times. ## Main definitions From 0723536a0522d24fc2f159a096fb3304bef77472 Mon Sep 17 00:00:00 2001 From: Riccardo Brasca Date: Mon, 19 Jun 2023 20:03:53 +0000 Subject: [PATCH 1201/1291] chore(data/zmod/algebra): make zmod.algebra a def (#19197) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit `zmod.algebra` creates a diamond about the `zmod p`-algebra structure on `zmod p`: ```lean import algebra.algebra.basic import data.zmod.algebra example (p : ℕ) : algebra.id (zmod p) = (zmod.algebra (zmod p) p) := rfl --fails ``` This is also causing troubles with the port. We turn `zmod.algebra` into a def. --- src/data/zmod/algebra.lean | 10 +++++++-- src/field_theory/cardinality.lean | 2 ++ src/field_theory/finite/galois_field.lean | 4 ++-- src/field_theory/finite/trace.lean | 7 ++++--- .../is_alg_closed/classification.lean | 21 ++++++++++--------- .../legendre_symbol/add_character.lean | 1 + .../polynomial/cyclotomic/expand.lean | 3 +++ src/ring_theory/witt_vector/frobenius.lean | 1 + 8 files changed, 32 insertions(+), 17 deletions(-) diff --git a/src/data/zmod/algebra.lean b/src/data/zmod/algebra.lean index 65f3f96fc884f..43e2ea791b9f9 100644 --- a/src/data/zmod/algebra.lean +++ b/src/data/zmod/algebra.lean @@ -24,7 +24,9 @@ instance (p : ℕ) : subsingleton (algebra (zmod p) R) := section variables {n : ℕ} (m : ℕ) [char_p R m] -/-- The `zmod n`-algebra structure on rings whose characteristic `m` divides `n` -/ +/-- The `zmod n`-algebra structure on rings whose characteristic `m` divides `n`. +See note [reducible non-instances]. -/ +@[reducible] def algebra' (h : m ∣ n) : algebra (zmod n) R := { smul := λ a r, a * r, commutes' := λ a r, show (a * r : R) = r * a, @@ -39,6 +41,10 @@ def algebra' (h : m ∣ n) : algebra (zmod n) R := end -instance (p : ℕ) [char_p R p] : algebra (zmod p) R := algebra' R p dvd_rfl +/-- The `zmod p`-algebra structure on a ring of characteristic `p`. This is not an +instance since it creates a diamond with `algebra.id`. +See note [reducible non-instances]. -/ +@[reducible] +def algebra (p : ℕ) [char_p R p] : algebra (zmod p) R := algebra' R p dvd_rfl end zmod diff --git a/src/field_theory/cardinality.lean b/src/field_theory/cardinality.lean index 9fe35f450678a..17a1c271f3e9b 100644 --- a/src/field_theory/cardinality.lean +++ b/src/field_theory/cardinality.lean @@ -36,8 +36,10 @@ universe u /-- A finite field has prime power cardinality. -/ lemma fintype.is_prime_pow_card_of_field {α} [fintype α] [field α] : is_prime_pow (‖α‖) := begin + -- TODO: `algebra` version of `char_p.exists`, of type `Σ p, algebra (zmod p) α` casesI char_p.exists α with p _, haveI hp := fact.mk (char_p.char_is_prime α p), + letI : algebra (zmod p) α := zmod.algebra _ _, let b := is_noetherian.finset_basis (zmod p) α, rw [module.card_fintype b, zmod.card, is_prime_pow_pow_iff], { exact hp.1.is_prime_pow }, diff --git a/src/field_theory/finite/galois_field.lean b/src/field_theory/finite/galois_field.lean index 8961eb821974a..8737cefa46b92 100644 --- a/src/field_theory/finite/galois_field.lean +++ b/src/field_theory/finite/galois_field.lean @@ -152,8 +152,6 @@ begin rw [splits_iff_card_roots, h1, ←finset.card_def, finset.card_univ, h2, zmod.card], end -local attribute [-instance] zmod.algebra - /-- A Galois field with exponent 1 is equivalent to `zmod` -/ def equiv_zmod_p : galois_field p 1 ≃ₐ[zmod p] (zmod p) := let h : (X ^ p ^ 1 : (zmod p)[X]) = X ^ (fintype.card (zmod p)) := @@ -229,6 +227,8 @@ begin all_goals {apply_instance}, }, rw ← hpp' at *, haveI := fact_iff.2 hp, + letI : algebra (zmod p) K := zmod.algebra _ _, + letI : algebra (zmod p) K' := zmod.algebra _ _, exact alg_equiv_of_card_eq p hKK', end diff --git a/src/field_theory/finite/trace.lean b/src/field_theory/finite/trace.lean index b5c3451ff78b4..42c334451a5f2 100644 --- a/src/field_theory/finite/trace.lean +++ b/src/field_theory/finite/trace.lean @@ -19,14 +19,15 @@ finite field, trace namespace finite_field /-- The trace map from a finite field to its prime field is nongedenerate. -/ -lemma trace_to_zmod_nondegenerate (F : Type*) [field F] [finite F] {a : F} - (ha : a ≠ 0) : ∃ b : F, algebra.trace (zmod (ring_char F)) F (a * b) ≠ 0 := +lemma trace_to_zmod_nondegenerate (F : Type*) [field F] [finite F] + [algebra (zmod (ring_char F)) F] {a : F} (ha : a ≠ 0) : + ∃ b : F, algebra.trace (zmod (ring_char F)) F (a * b) ≠ 0 := begin haveI : fact (ring_char F).prime := ⟨char_p.char_is_prime F _⟩, have htr := trace_form_nondegenerate (zmod (ring_char F)) F a, simp_rw [algebra.trace_form_apply] at htr, by_contra' hf, - exact ha (htr hf), + exact ha (htr hf) end end finite_field diff --git a/src/field_theory/is_alg_closed/classification.lean b/src/field_theory/is_alg_closed/classification.lean index 169081196a4db..77d72392bc026 100644 --- a/src/field_theory/is_alg_closed/classification.lean +++ b/src/field_theory/is_alg_closed/classification.lean @@ -158,10 +158,9 @@ variables {K L : Type} [field K] [field L] [is_alg_closed K] [is_alg_closed L] /-- Two uncountable algebraically closed fields of characteristic zero are isomorphic if they have the same cardinality. -/ -@[nolint def_lemma] lemma ring_equiv_of_cardinal_eq_of_char_zero [char_zero K] [char_zero L] - (hK : ℵ₀ < #K) (hKL : #K = #L) : K ≃+* L := +lemma ring_equiv_of_cardinal_eq_of_char_zero [char_zero K] [char_zero L] + (hK : ℵ₀ < #K) (hKL : #K = #L) : nonempty (K ≃+* L) := begin - apply classical.choice, cases exists_is_transcendence_basis ℤ (show function.injective (algebra_map ℤ K), from int.cast_injective) with s hs, @@ -177,9 +176,10 @@ begin end private lemma ring_equiv_of_cardinal_eq_of_char_p (p : ℕ) [fact p.prime] - [char_p K p] [char_p L p] (hK : ℵ₀ < #K) (hKL : #K = #L) : K ≃+* L := + [char_p K p] [char_p L p] (hK : ℵ₀ < #K) (hKL : #K = #L) : nonempty (K ≃+* L) := begin - apply classical.choice, + letI : algebra (zmod p) K := zmod.algebra _ _, + letI : algebra (zmod p) L := zmod.algebra _ _, cases exists_is_transcendence_basis (zmod p) (show function.injective (algebra_map (zmod p) K), from ring_hom.injective _) with s hs, @@ -198,18 +198,19 @@ end /-- Two uncountable algebraically closed fields are isomorphic if they have the same cardinality and the same characteristic. -/ -@[nolint def_lemma] lemma ring_equiv_of_cardinal_eq_of_char_eq (p : ℕ) [char_p K p] [char_p L p] - (hK : ℵ₀ < #K) (hKL : #K = #L) : K ≃+* L := +lemma ring_equiv_of_cardinal_eq_of_char_eq (p : ℕ) [char_p K p] [char_p L p] + (hK : ℵ₀ < #K) (hKL : #K = #L) : nonempty (K ≃+* L) := begin - apply classical.choice, rcases char_p.char_is_prime_or_zero K p with hp | hp, { haveI : fact p.prime := ⟨hp⟩, - exact ⟨ring_equiv_of_cardinal_eq_of_char_p p hK hKL⟩ }, + letI : algebra (zmod p) K := zmod.algebra _ _, + letI : algebra (zmod p) L := zmod.algebra _ _, + exact ring_equiv_of_cardinal_eq_of_char_p p hK hKL }, { rw [hp] at *, resetI, letI : char_zero K := char_p.char_p_to_char_zero K, letI : char_zero L := char_p.char_p_to_char_zero L, - exact ⟨ring_equiv_of_cardinal_eq_of_char_zero hK hKL⟩ } + exact ring_equiv_of_cardinal_eq_of_char_zero hK hKL } end end is_alg_closed diff --git a/src/number_theory/legendre_symbol/add_character.lean b/src/number_theory/legendre_symbol/add_character.lean index 94407f97e1d82..e35db8da04b63 100644 --- a/src/number_theory/legendre_symbol/add_character.lean +++ b/src/number_theory/legendre_symbol/add_character.lean @@ -355,6 +355,7 @@ begin exact λ hf, nat.prime.ne_zero hp.1 (zero_dvd_iff.mp hf), }, end, let ψ := primitive_zmod_char pp F' (ne_zero_iff.mp (ne_zero.of_not_dvd F' hp₂)), + letI : algebra (zmod p) F := zmod.algebra _ _, let ψ' := ψ.char.comp (algebra.trace (zmod p) F).to_add_monoid_hom.to_multiplicative, have hψ' : is_nontrivial ψ' := begin diff --git a/src/ring_theory/polynomial/cyclotomic/expand.lean b/src/ring_theory/polynomial/cyclotomic/expand.lean index f79fa27ec789a..72a5e23858061 100644 --- a/src/ring_theory/polynomial/cyclotomic/expand.lean +++ b/src/ring_theory/polynomial/cyclotomic/expand.lean @@ -126,6 +126,7 @@ section char_p lemma cyclotomic_mul_prime_eq_pow_of_not_dvd (R : Type*) {p n : ℕ} [hp : fact (nat.prime p)] [ring R] [char_p R p] (hn : ¬p ∣ n) : cyclotomic (n * p) R = (cyclotomic n R) ^ (p - 1) := begin + letI : algebra (zmod p) R := zmod.algebra _ _, suffices : cyclotomic (n * p) (zmod p) = (cyclotomic n (zmod p)) ^ (p - 1), { rw [← map_cyclotomic _ (algebra_map (zmod p) R), ← map_cyclotomic _ (algebra_map (zmod p) R), this, polynomial.map_pow] }, @@ -141,6 +142,7 @@ end lemma cyclotomic_mul_prime_dvd_eq_pow (R : Type*) {p n : ℕ} [hp : fact (nat.prime p)] [ring R] [char_p R p] (hn : p ∣ n) : cyclotomic (n * p) R = (cyclotomic n R) ^ p := begin + letI : algebra (zmod p) R := zmod.algebra _ _, suffices : cyclotomic (n * p) (zmod p) = (cyclotomic n (zmod p)) ^ p, { rw [← map_cyclotomic _ (algebra_map (zmod p) R), ← map_cyclotomic _ (algebra_map (zmod p) R), this, polynomial.map_pow] }, @@ -171,6 +173,7 @@ lemma is_root_cyclotomic_prime_pow_mul_iff_of_char_p {m k p : ℕ} {R : Type*} [ [is_domain R] [hp : fact (nat.prime p)] [hchar : char_p R p] {μ : R} [ne_zero (m : R)] : (polynomial.cyclotomic (p ^ k * m) R).is_root μ ↔ is_primitive_root μ m := begin + letI : algebra (zmod p) R := zmod.algebra _ _, rcases k.eq_zero_or_pos with rfl | hk, { rw [pow_zero, one_mul, is_root_cyclotomic_iff] }, refine ⟨λ h, _, λ h, _⟩, diff --git a/src/ring_theory/witt_vector/frobenius.lean b/src/ring_theory/witt_vector/frobenius.lean index a5b97d84730a9..912a4b5962c9a 100644 --- a/src/ring_theory/witt_vector/frobenius.lean +++ b/src/ring_theory/witt_vector/frobenius.lean @@ -282,6 +282,7 @@ lemma coeff_frobenius_char_p (x : 𝕎 R) (n : ℕ) : coeff (frobenius x) n = (x.coeff n) ^ p := begin rw [coeff_frobenius], + letI : algebra (zmod p) R := zmod.algebra _ _, -- outline of the calculation, proofs follow below calc aeval (λ k, x.coeff k) (frobenius_poly p n) = aeval (λ k, x.coeff k) From 86c29aefdba50b3f33e86e52e3b2f51a0d8f0282 Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Tue, 20 Jun 2023 04:52:54 +0000 Subject: [PATCH 1202/1291] refactor(geometry/manifold/cont_mdiff_map): redefine as subtype not structure (#19147) We have a long-running conversation in mathlib about whether function spaces should be implemented as subtypes or as structures. This PR proposes to change `cont_mdiff_map`, the type of smooth functions between manifolds `M` and `M'`, from a "structure" implementation to a "subtype" implementation. It honestly seems pretty painless, even though this is a widely used type -- the only change for users is that the field names are now `val` and `property` rather than `to_fun` and `cont_mdiff_to_fun`. The motivation is to make it possible to make certain constructions about function spaces generic, so that work for the space of smooth functions can be reused for (for example) the spaces of continuous or differentiable functions. Notably, in #19146 we introduce a generic construction of a sheaf of functions on a manifold whose object over the open set `U` is the subtype of functions satisfying a "local invariant property". With this PR, when that construction is applied to the property "smoothness", the resulting sheaf has objects which are *by definition* the types `cont_mdiff_map`. They then inherit algebraic structures for free, see #19094. --- src/geometry/manifold/cont_mdiff_map.lean | 32 ++++++++++---------- src/geometry/manifold/derivation_bundle.lean | 4 +-- src/geometry/manifold/whitney_embedding.lean | 4 +-- 3 files changed, 20 insertions(+), 20 deletions(-) diff --git a/src/geometry/manifold/cont_mdiff_map.lean b/src/geometry/manifold/cont_mdiff_map.lean index 1fce5586befc0..9b736749c8f52 100644 --- a/src/geometry/manifold/cont_mdiff_map.lean +++ b/src/geometry/manifold/cont_mdiff_map.lean @@ -33,10 +33,7 @@ variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] (n : ℕ∞) /-- Bundled `n` times continuously differentiable maps. -/ -@[protect_proj] -structure cont_mdiff_map := -(to_fun : M → M') -(cont_mdiff_to_fun : cont_mdiff I I' n to_fun) +def cont_mdiff_map := {f : M → M' // cont_mdiff I I' n f} /-- Bundled smooth maps. -/ @[reducible] def smooth_map := cont_mdiff_map I I' M M' ⊤ @@ -52,24 +49,27 @@ namespace cont_mdiff_map variables {I} {I'} {M} {M'} {n} -instance : has_coe_to_fun C^n⟮I, M; I', M'⟯ (λ _, M → M') := ⟨cont_mdiff_map.to_fun⟩ +instance fun_like : fun_like C^n⟮I, M; I', M'⟯ M (λ _, M') := +{ coe := subtype.val, + coe_injective' := subtype.coe_injective } + +protected lemma cont_mdiff (f : C^n⟮I, M; I', M'⟯) : + cont_mdiff I I' n f := f.prop + +protected lemma smooth (f : C^∞⟮I, M; I', M'⟯) : + smooth I I' f := f.prop + instance : has_coe C^n⟮I, M; I', M'⟯ C(M, M') := -⟨λ f, ⟨f, f.cont_mdiff_to_fun.continuous⟩⟩ +⟨λ f, ⟨f, f.cont_mdiff.continuous⟩⟩ attribute [to_additive_ignore_args 21] cont_mdiff_map - cont_mdiff_map.has_coe_to_fun cont_mdiff_map.continuous_map.has_coe + cont_mdiff_map.fun_like cont_mdiff_map.continuous_map.has_coe variables {f g : C^n⟮I, M; I', M'⟯} @[simp] lemma coe_fn_mk (f : M → M') (hf : cont_mdiff I I' n f) : - (mk f hf : M → M') = f := + ((by exact subtype.mk f hf : C^n⟮I, M; I', M'⟯) : M → M') = f := rfl -protected lemma cont_mdiff (f : C^n⟮I, M; I', M'⟯) : - cont_mdiff I I' n f := f.cont_mdiff_to_fun - -protected lemma smooth (f : C^∞⟮I, M; I', M'⟯) : - smooth I I' f := f.cont_mdiff_to_fun - lemma coe_inj ⦃f g : C^n⟮I, M; I', M'⟯⦄ (h : (f : M → M') = g) : f = g := by cases f; cases g; cases h; refl @@ -86,8 +86,8 @@ def id : C^n⟮I, M; I, M⟯ := ⟨id, cont_mdiff_id⟩ /-- The composition of smooth maps, as a smooth map. -/ def comp (f : C^n⟮I', M'; I'', M''⟯) (g : C^n⟮I, M; I', M'⟯) : C^n⟮I, M; I'', M''⟯ := -{ to_fun := λ a, f (g a), - cont_mdiff_to_fun := f.cont_mdiff_to_fun.comp g.cont_mdiff_to_fun, } +{ val := λ a, f (g a), + property := f.cont_mdiff.comp g.cont_mdiff, } @[simp] lemma comp_apply (f : C^n⟮I', M'; I'', M''⟯) (g : C^n⟮I, M; I', M'⟯) (x : M) : f.comp g x = f (g x) := rfl diff --git a/src/geometry/manifold/derivation_bundle.lean b/src/geometry/manifold/derivation_bundle.lean index 7a37e9acc559b..71f103747b268 100644 --- a/src/geometry/manifold/derivation_bundle.lean +++ b/src/geometry/manifold/derivation_bundle.lean @@ -42,8 +42,8 @@ variables {𝕜 M} namespace pointed_smooth_map -instance {x : M} : has_coe_to_fun C^∞⟮I, M; 𝕜⟯⟨x⟩ (λ _, M → 𝕜) := -cont_mdiff_map.has_coe_to_fun +instance fun_like {x : M} : fun_like C^∞⟮I, M; 𝕜⟯⟨x⟩ M (λ _, 𝕜) := +cont_mdiff_map.fun_like instance {x : M} : comm_ring C^∞⟮I, M; 𝕜⟯⟨x⟩ := smooth_map.comm_ring instance {x : M} : algebra 𝕜 C^∞⟮I, M; 𝕜⟯⟨x⟩ := smooth_map.algebra instance {x : M} : inhabited C^∞⟮I, M; 𝕜⟯⟨x⟩ := ⟨0⟩ diff --git a/src/geometry/manifold/whitney_embedding.lean b/src/geometry/manifold/whitney_embedding.lean index 02240f02a1f1a..736ba534e91f6 100644 --- a/src/geometry/manifold/whitney_embedding.lean +++ b/src/geometry/manifold/whitney_embedding.lean @@ -50,8 +50,8 @@ include hi /-- Smooth embedding of `M` into `(E × ℝ) ^ ι`. -/ def embedding_pi_tangent : C^∞⟮I, M; 𝓘(ℝ, ι → (E × ℝ)), ι → (E × ℝ)⟯ := -{ to_fun := λ x i, (f i x • ext_chart_at I (f.c i) x, f i x), - cont_mdiff_to_fun := cont_mdiff_pi_space.2 $ λ i, +{ val := λ x i, (f i x • ext_chart_at I (f.c i) x, f i x), + property := cont_mdiff_pi_space.2 $ λ i, ((f i).smooth_smul cont_mdiff_on_ext_chart_at).prod_mk_space ((f i).smooth) } local attribute [simp] lemma embedding_pi_tangent_coe : From f2ad3645af9effcdb587637dc28a6074edc813f9 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Tue, 20 Jun 2023 06:00:43 +0000 Subject: [PATCH 1203/1291] chore(*): add mathlib4 synchronization comments (#19207) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebraic_geometry.open_immersion.basic` * `analysis.schwartz_space` * `category_theory.monoidal.internal.types` * `dynamics.ergodic.add_circle` * `field_theory.is_alg_closed.algebraic_closure` * `measure_theory.group.add_circle` * `measure_theory.integral.periodic` * `number_theory.kummer_dedekind` * `number_theory.well_approximable` * `probability.kernel.integral_comp_prod` * `probability.martingale.basic` * `probability.martingale.centering` * `probability.martingale.optional_sampling` * `probability.process.hitting_time` * `ring_theory.is_adjoin_root` --- src/algebraic_geometry/open_immersion/basic.lean | 3 +++ src/analysis/schwartz_space.lean | 3 +++ src/category_theory/monoidal/internal/types.lean | 3 +++ src/dynamics/ergodic/add_circle.lean | 3 +++ src/field_theory/is_alg_closed/algebraic_closure.lean | 3 +++ src/measure_theory/group/add_circle.lean | 3 +++ src/measure_theory/integral/periodic.lean | 3 +++ src/number_theory/kummer_dedekind.lean | 3 +++ src/number_theory/well_approximable.lean | 3 +++ src/probability/kernel/integral_comp_prod.lean | 3 +++ src/probability/martingale/basic.lean | 3 +++ src/probability/martingale/centering.lean | 3 +++ src/probability/martingale/optional_sampling.lean | 3 +++ src/probability/process/hitting_time.lean | 3 +++ src/ring_theory/is_adjoin_root.lean | 3 +++ 15 files changed, 45 insertions(+) diff --git a/src/algebraic_geometry/open_immersion/basic.lean b/src/algebraic_geometry/open_immersion/basic.lean index c4c7b15bb57de..0592e7c1bed51 100644 --- a/src/algebraic_geometry/open_immersion/basic.lean +++ b/src/algebraic_geometry/open_immersion/basic.lean @@ -9,6 +9,9 @@ import algebraic_geometry.locally_ringed_space /-! # Open immersions of structured spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We say that a morphism of presheafed spaces `f : X ⟶ Y` is an open immersion if the underlying map of spaces is an open embedding `f : X ⟶ U ⊆ Y`, and the sheaf map `Y(V) ⟶ f _* X(V)` is an iso for each `V ⊆ U`. diff --git a/src/analysis/schwartz_space.lean b/src/analysis/schwartz_space.lean index 085436610ee0f..86c4d8105c34e 100644 --- a/src/analysis/schwartz_space.lean +++ b/src/analysis/schwartz_space.lean @@ -15,6 +15,9 @@ import analysis.special_functions.pow.real /-! # Schwartz space +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the Schwartz space. Usually, the Schwartz space is defined as the set of smooth functions $f : ℝ^n → ℂ$ such that there exists $C_{αβ} > 0$ with $$|x^α ∂^β f(x)| < C_{αβ}$$ for all $x ∈ ℝ^n$ and for all multiindices $α, β$. diff --git a/src/category_theory/monoidal/internal/types.lean b/src/category_theory/monoidal/internal/types.lean index 6351e83a93dac..39831cf76281a 100644 --- a/src/category_theory/monoidal/internal/types.lean +++ b/src/category_theory/monoidal/internal/types.lean @@ -10,6 +10,9 @@ import category_theory.monoidal.types.symmetric /-! # `Mon_ (Type u) ≌ Mon.{u}` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The category of internal monoid objects in `Type` is equivalent to the category of "native" bundled monoids. diff --git a/src/dynamics/ergodic/add_circle.lean b/src/dynamics/ergodic/add_circle.lean index 18955e864bfe8..8bac8c19a8a8b 100644 --- a/src/dynamics/ergodic/add_circle.lean +++ b/src/dynamics/ergodic/add_circle.lean @@ -11,6 +11,9 @@ import data.set.pointwise.iterate /-! # Ergodic maps of the additive circle +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains proofs of ergodicity for maps of the additive circle. ## Main definitions: diff --git a/src/field_theory/is_alg_closed/algebraic_closure.lean b/src/field_theory/is_alg_closed/algebraic_closure.lean index 5f7ead381cc16..5aa7ca14f72f6 100644 --- a/src/field_theory/is_alg_closed/algebraic_closure.lean +++ b/src/field_theory/is_alg_closed/algebraic_closure.lean @@ -11,6 +11,9 @@ import field_theory.splitting_field.construction /-! # Algebraic Closure +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we construct the algebraic closure of a field ## Main Definitions diff --git a/src/measure_theory/group/add_circle.lean b/src/measure_theory/group/add_circle.lean index 50e9ffcab8e9c..67d2fc9588ddf 100644 --- a/src/measure_theory/group/add_circle.lean +++ b/src/measure_theory/group/add_circle.lean @@ -9,6 +9,9 @@ import data.zmod.quotient /-! # Measure-theoretic results about the additive circle +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The file is a place to collect measure-theoretic results about the additive circle. ## Main definitions: diff --git a/src/measure_theory/integral/periodic.lean b/src/measure_theory/integral/periodic.lean index f34bc3b914757..490cba056c617 100644 --- a/src/measure_theory/integral/periodic.lean +++ b/src/measure_theory/integral/periodic.lean @@ -12,6 +12,9 @@ import topology.algebra.order.floor /-! # Integrals of periodic functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove that the half-open interval `Ioc t (t + T)` in `ℝ` is a fundamental domain of the action of the subgroup `ℤ ∙ T` on `ℝ`. diff --git a/src/number_theory/kummer_dedekind.lean b/src/number_theory/kummer_dedekind.lean index 0ee9f05a5c58d..28af342219f88 100644 --- a/src/number_theory/kummer_dedekind.lean +++ b/src/number_theory/kummer_dedekind.lean @@ -10,6 +10,9 @@ import ring_theory.is_adjoin_root /-! # Kummer-Dedekind theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves the monogenic version of the Kummer-Dedekind theorem on the splitting of prime ideals in an extension of the ring of integers. This states that if `I` is a prime ideal of Dedekind domain `R` and `S = R[α]` for some `α` that is integral over `R` with minimal polynomial diff --git a/src/number_theory/well_approximable.lean b/src/number_theory/well_approximable.lean index 93e703e48f198..dc65fb61b0c4e 100644 --- a/src/number_theory/well_approximable.lean +++ b/src/number_theory/well_approximable.lean @@ -9,6 +9,9 @@ import measure_theory.covering.liminf_limsup /-! # Well-approximable numbers and Gallagher's ergodic theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Gallagher's ergodic theorem is a result in metric number theory. It thus belongs to that branch of mathematics concerning arithmetic properties of real numbers which hold almost eveywhere with respect to the Lebesgue measure. diff --git a/src/probability/kernel/integral_comp_prod.lean b/src/probability/kernel/integral_comp_prod.lean index 7b6917274b16f..53b9df2f7c6e8 100644 --- a/src/probability/kernel/integral_comp_prod.lean +++ b/src/probability/kernel/integral_comp_prod.lean @@ -9,6 +9,9 @@ import measure_theory.integral.set_integral /-! # Bochner integral of a function against the composition-product of two kernels +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We prove properties of the composition-product of two kernels. If `κ` is an s-finite kernel from `α` to `β` and `η` is an s-finite kernel from `α × β` to `γ`, we can form their composition-product `κ ⊗ₖ η : kernel α (β × γ)`. We proved in `probability.kernel.lintegral_comp_prod` that it verifies diff --git a/src/probability/martingale/basic.lean b/src/probability/martingale/basic.lean index 9501a4443e29f..e4cb7613f4d6f 100644 --- a/src/probability/martingale/basic.lean +++ b/src/probability/martingale/basic.lean @@ -9,6 +9,9 @@ import probability.process.stopping /-! # Martingales +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A family of functions `f : ι → Ω → E` is a martingale with respect to a filtration `ℱ` if every `f i` is integrable, `f` is adapted with respect to `ℱ` and for all `i ≤ j`, `μ[f j | ℱ i] =ᵐ[μ] f i`. On the other hand, `f : ι → Ω → E` is said to be a supermartingale diff --git a/src/probability/martingale/centering.lean b/src/probability/martingale/centering.lean index f5ddff3bd5b35..06bbc68ec6a4a 100644 --- a/src/probability/martingale/centering.lean +++ b/src/probability/martingale/centering.lean @@ -9,6 +9,9 @@ import probability.martingale.basic /-! # Centering lemma for stochastic processes +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Any `ℕ`-indexed stochastic process which is adapted and integrable can be written as the sum of a martingale and a predictable process. This result is also known as **Doob's decomposition theorem**. From a process `f`, a filtration `ℱ` and a measure `μ`, we define two processes diff --git a/src/probability/martingale/optional_sampling.lean b/src/probability/martingale/optional_sampling.lean index 61092ba02e0a4..ac1c91b34c817 100644 --- a/src/probability/martingale/optional_sampling.lean +++ b/src/probability/martingale/optional_sampling.lean @@ -10,6 +10,9 @@ import probability.martingale.basic /-! # Optional sampling theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + If `τ` is a bounded stopping time and `σ` is another stopping time, then the value of a martingale `f` at the stopping time `min τ σ` is almost everywhere equal to `μ[stopped_value f τ | hσ.measurable_space]`. diff --git a/src/probability/process/hitting_time.lean b/src/probability/process/hitting_time.lean index 7881a97d1152d..f8260177b6912 100644 --- a/src/probability/process/hitting_time.lean +++ b/src/probability/process/hitting_time.lean @@ -8,6 +8,9 @@ import probability.process.stopping /-! # Hitting time +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given a stochastic process, the hitting time provides the first time the process ``hits'' some subset of the state space. The hitting time is a stopping time in the case that the time index is discrete and the process is adapted (this is true in a far more general setting however we have diff --git a/src/ring_theory/is_adjoin_root.lean b/src/ring_theory/is_adjoin_root.lean index 087a07edd462b..b13d9a322511e 100644 --- a/src/ring_theory/is_adjoin_root.lean +++ b/src/ring_theory/is_adjoin_root.lean @@ -10,6 +10,9 @@ import ring_theory.power_basis /-! # A predicate on adjoining roots of polynomial +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines a predicate `is_adjoin_root S f`, which states that the ring `S` can be constructed by adjoining a specified root of the polynomial `f : R[X]` to `R`. This predicate is useful when the same ring can be generated by adjoining the root of different From 431589bce478b2229eba14b14a283250428217db Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Tue, 20 Jun 2023 21:31:46 +0000 Subject: [PATCH 1204/1291] feat(geometry/manifold/sheaf/basic): sheaf of functions satisfying a `local_invariant_prop` (#19146) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Define `structure_groupoid.local_invariant_prop.sheaf`, the sheaf-of-types of functions `f : M → M'` (for charted spaces `M`, `M'`) satisfying some local property in the sense of `structure_groupoid.local_invariant_prop` (for example continuity, differentiability, smoothness). --- src/geometry/manifold/charted_space.lean | 15 ++++ .../manifold/local_invariant_properties.lean | 24 ++++- src/geometry/manifold/sheaf/basic.lean | 89 +++++++++++++++++++ src/topology/local_homeomorph.lean | 31 +++++++ 4 files changed, 158 insertions(+), 1 deletion(-) create mode 100644 src/geometry/manifold/sheaf/basic.lean diff --git a/src/geometry/manifold/charted_space.lean b/src/geometry/manifold/charted_space.lean index 519ab352e5316..9d562746d5252 100644 --- a/src/geometry/manifold/charted_space.lean +++ b/src/geometry/manifold/charted_space.lean @@ -984,6 +984,21 @@ instance [closed_under_restriction G] : has_groupoid s G := { exact preimage_open_of_open_symm (chart_at H x) s.2 }, end } +lemma chart_at_inclusion_symm_eventually_eq {U V : opens M} (hUV : U ≤ V) {x : U} : + (chart_at H (set.inclusion hUV x)).symm + =ᶠ[𝓝 (chart_at H (set.inclusion hUV x) (set.inclusion hUV x))] set.inclusion hUV + ∘ (chart_at H x).symm := +begin + set i := set.inclusion hUV, + set e := chart_at H (x:M), + haveI : nonempty U := ⟨x⟩, + haveI : nonempty V := ⟨i x⟩, + have heUx_nhds : (e.subtype_restr U).target ∈ 𝓝 (e x), + { apply (e.subtype_restr U).open_target.mem_nhds, + exact e.map_subtype_source (mem_chart_source _ _) }, + exact filter.eventually_eq_of_mem heUx_nhds (e.subtype_restr_symm_eq_on_of_le hUV), +end + end topological_space.opens /-! ### Structomorphisms -/ diff --git a/src/geometry/manifold/local_invariant_properties.lean b/src/geometry/manifold/local_invariant_properties.lean index 8e07981ff9550..fedd470ff9b59 100644 --- a/src/geometry/manifold/local_invariant_properties.lean +++ b/src/geometry/manifold/local_invariant_properties.lean @@ -48,7 +48,7 @@ in the one for `lift_prop_within_at`. noncomputable theory open_locale classical manifold topology -open set filter +open set filter topological_space variables {H M H' M' X : Type*} variables [topological_space H] [topological_space M] [charted_space H M] @@ -569,6 +569,28 @@ begin exact λ x, hG.congr' ((chart_at H x).eventually_right_inverse $ mem_chart_target H x) (hQ _) end +lemma lift_prop_at_iff_comp_inclusion (hG : local_invariant_prop G G' P) {U V : opens M} + (hUV : U ≤ V) (f : V → M') (x : U) : + lift_prop_at P f (set.inclusion hUV x) ↔ lift_prop_at P (f ∘ set.inclusion hUV : U → M') x := +begin + congrm _ ∧ _, + { simp [continuous_within_at_univ, + (topological_space.opens.open_embedding_of_le hUV).continuous_at_iff] }, + { apply hG.congr_iff, + exact (topological_space.opens.chart_at_inclusion_symm_eventually_eq hUV).fun_comp + (chart_at H' (f (set.inclusion hUV x)) ∘ f) }, +end + +lemma lift_prop_inclusion {Q : (H → H) → (set H) → H → Prop} (hG : local_invariant_prop G G Q) + (hQ : ∀ y, Q id univ y) {U V : opens M} (hUV : U ≤ V) : + lift_prop Q (set.inclusion hUV : U → V) := +begin + intro x, + show lift_prop_at Q (id ∘ inclusion hUV) x, + rw ← hG.lift_prop_at_iff_comp_inclusion hUV, + apply hG.lift_prop_id hQ, +end + end local_invariant_prop section local_structomorph diff --git a/src/geometry/manifold/sheaf/basic.lean b/src/geometry/manifold/sheaf/basic.lean new file mode 100644 index 0000000000000..35fe4545d0d01 --- /dev/null +++ b/src/geometry/manifold/sheaf/basic.lean @@ -0,0 +1,89 @@ +/- +Copyright © 2023 Heather Macbeth. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Heather Macbeth +-/ +import geometry.manifold.local_invariant_properties +import topology.sheaves.local_predicate + +/-! # Generic construction of a sheaf from a `local_invariant_prop` on a manifold + +This file constructs the sheaf-of-types of functions `f : M → M'` (for charted spaces `M`, `M'`) +which satisfy the lifted property `lift_prop P` associated to some locally invariant (in the sense +of `structure_groupoid.local_invariant_prop`) property `P` on the model spaces of `M` and `M'`. For +example, differentiability and smoothness are locally invariant properties in this sense, so this +construction can be used to construct the sheaf of differentiable functions on a manifold and the +sheaf of smooth functions on a manifold. + +The mathematical work is in associating a `Top.local_predicate` to a +`structure_groupoid.local_invariant_prop`: that is, showing that a differential-geometric "locally +invariant" property is preserved under restriction and gluing. + +## Main definitions + +* `structure_groupoid.local_invariant_prop.local_predicate`: the `Top.local_predicate` (in the + sheaf-theoretic sense) on functions from open subsets of `M` into `M'`, which states whether + such functions satisfy `lift_prop P`. +* `structure_groupoid.local_invariant_prop.sheaf`: the sheaf-of-types of functions `f : M → M'` + which satisfy the lifted property `lift_prop P`. +-/ + +open_locale manifold topology +open set topological_space structure_groupoid structure_groupoid.local_invariant_prop opposite + +universe u + +variables {H : Type*} [topological_space H] {H' : Type*} [topological_space H'] + {G : structure_groupoid H} {G' : structure_groupoid H'} + {P : (H → H') → (set H) → H → Prop} + (M : Type u) [topological_space M] [charted_space H M] + (M' : Type u) [topological_space M'] [charted_space H' M'] + +instance Top.of.charted_space : charted_space H (Top.of M) := (infer_instance : charted_space H M) + +instance Top.of.has_groupoid [has_groupoid M G] : has_groupoid (Top.of M) G := +(infer_instance : has_groupoid M G) + +/-- Let `P` be a `local_invariant_prop` for functions between spaces with the groupoids `G`, `G'` +and let `M`, `M'` be charted spaces modelled on the model spaces of those groupoids. Then there is +an induced `local_predicate` on the functions from `M` to `M'`, given by `lift_prop P`. -/ +def structure_groupoid.local_invariant_prop.local_predicate (hG : local_invariant_prop G G' P) : + Top.local_predicate (λ (x : Top.of M), M') := +{ pred := λ {U : opens (Top.of M)}, λ (f : U → M'), lift_prop P f, + res := begin + intros U V i f h x, + have hUV : U ≤ V := category_theory.le_of_hom i, + show lift_prop_at P (f ∘ set.inclusion hUV) x, + rw ← hG.lift_prop_at_iff_comp_inclusion hUV, + apply h, + end, + locality := begin + intros V f h x, + obtain ⟨U, hxU, i, hU : lift_prop P (f ∘ i)⟩ := h x, + let x' : U := ⟨x, hxU⟩, + have hUV : U ≤ V := category_theory.le_of_hom i, + have : lift_prop_at P f (inclusion hUV x'), + { rw hG.lift_prop_at_iff_comp_inclusion hUV, + exact hU x' }, + convert this, + ext1, + refl + end } + +/-- Let `P` be a `local_invariant_prop` for functions between spaces with the groupoids `G`, `G'` +and let `M`, `M'` be charted spaces modelled on the model spaces of those groupoids. Then there is +a sheaf of types on `M` which, to each open set `U` in `M`, associates the type of bundled +functions from `U` to `M'` satisfying the lift of `P`. -/ +def structure_groupoid.local_invariant_prop.sheaf (hG : local_invariant_prop G G' P) : + Top.sheaf (Type u) (Top.of M) := +Top.subsheaf_to_Types (hG.local_predicate M M') + +instance structure_groupoid.local_invariant_prop.sheaf_has_coe_to_fun + (hG : local_invariant_prop G G' P) (U : (opens (Top.of M))ᵒᵖ) : + has_coe_to_fun ((hG.sheaf M M').val.obj U) (λ _, (unop U) → M') := +{ coe := λ a, a.1 } + +lemma structure_groupoid.local_invariant_prop.section_spec (hG : local_invariant_prop G G' P) + (U : (opens (Top.of M))ᵒᵖ) (f : (hG.sheaf M M').val.obj U) : + lift_prop P f := +f.2 diff --git a/src/topology/local_homeomorph.lean b/src/topology/local_homeomorph.lean index 64ab36280df56..0e632d97b7f51 100644 --- a/src/topology/local_homeomorph.lean +++ b/src/topology/local_homeomorph.lean @@ -1146,6 +1146,17 @@ lemma subtype_restr_def : e.subtype_restr s = s.local_homeomorph_subtype_coe.tra @[simp, mfld_simps] lemma subtype_restr_source : (e.subtype_restr s).source = coe ⁻¹' e.source := by simp only [subtype_restr_def] with mfld_simps +variables {s} + +lemma map_subtype_source {x : s} (hxe : (x:α) ∈ e.source) : e x ∈ (e.subtype_restr s).target := +begin + refine ⟨e.map_source hxe, _⟩, + rw [s.local_homeomorph_subtype_coe_target, mem_preimage, e.left_inv_on hxe], + exact x.prop +end + +variables (s) + /- This lemma characterizes the transition functions of an open subset in terms of the transition functions of the original space. -/ lemma subtype_restr_symm_trans_subtype_restr (f f' : local_homeomorph α β) : @@ -1167,4 +1178,24 @@ begin simp only with mfld_simps, end +lemma subtype_restr_symm_eq_on_of_le {U V : opens α} [nonempty U] [nonempty V] (hUV : U ≤ V) : + eq_on (e.subtype_restr V).symm (set.inclusion hUV ∘ (e.subtype_restr U).symm) + (e.subtype_restr U).target := +begin + set i := set.inclusion hUV, + intros y hy, + dsimp [local_homeomorph.subtype_restr_def] at ⊢ hy, + have hyV : e.symm y ∈ V.local_homeomorph_subtype_coe.target, + { rw opens.local_homeomorph_subtype_coe_target at ⊢ hy, + exact hUV hy.2 }, + refine V.local_homeomorph_subtype_coe.inj_on _ trivial _, + { rw ←local_homeomorph.symm_target, + apply local_homeomorph.map_source, + rw local_homeomorph.symm_source, + exact hyV }, + { rw V.local_homeomorph_subtype_coe.right_inv hyV, + show _ = U.local_homeomorph_subtype_coe _, + rw U.local_homeomorph_subtype_coe.right_inv hy.2 } +end + end local_homeomorph From 722b3b152ddd5e0cf21c0a29787c76596cb6b422 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 20 Jun 2023 23:14:57 +0000 Subject: [PATCH 1205/1291] refactor(data/matrix/invertible): more results about invertible matrices (#19204) Many results about `invertible` apply directly to matrices simply by replacing `*` with `matrix.mul`. This also adds some missing lemmas about invertibility of products. --- src/algebra/invertible.lean | 55 ++++++++++- src/data/matrix/invertible.lean | 91 +++++++++++++++++++ .../matrix/nonsingular_inverse.lean | 26 +----- 3 files changed, 144 insertions(+), 28 deletions(-) create mode 100644 src/data/matrix/invertible.lean diff --git a/src/algebra/invertible.lean b/src/algebra/invertible.lean index 6728fe56ad325..9e64de20e92cb 100644 --- a/src/algebra/invertible.lean +++ b/src/algebra/invertible.lean @@ -109,12 +109,19 @@ by { apply inv_of_eq_right_inv, rw [h, mul_inv_of_self], } instance [monoid α] (a : α) : subsingleton (invertible a) := ⟨ λ ⟨b, hba, hab⟩ ⟨c, hca, hac⟩, by { congr, exact left_inv_eq_right_inv hba hac } ⟩ +/-- If `r` is invertible and `s = r` and `si = ⅟r`, then `s` is invertible with `⅟s = si`. -/ +def invertible.copy' [mul_one_class α] {r : α} (hr : invertible r) (s : α) (si : α) + (hs : s = r) (hsi : si = ⅟r) : + invertible s := +{ inv_of := si, + inv_of_mul_self := by rw [hs, hsi, inv_of_mul_self], + mul_inv_of_self := by rw [hs, hsi, mul_inv_of_self] } + /-- If `r` is invertible and `s = r`, then `s` is invertible. -/ +@[reducible] def invertible.copy [mul_one_class α] {r : α} (hr : invertible r) (s : α) (hs : s = r) : invertible s := -{ inv_of := ⅟r, - inv_of_mul_self := by rw [hs, inv_of_mul_self], - mul_inv_of_self := by rw [hs, mul_inv_of_self] } +hr.copy' _ _ hs rfl /-- An `invertible` element is a unit. -/ @[simps] @@ -196,6 +203,11 @@ def invertible_mul [monoid α] (a b : α) [invertible a] [invertible b] : invert ⅟(a * b) = ⅟b * ⅟a := inv_of_eq_right_inv (by simp [←mul_assoc]) +/-- A copy of `invertible_mul` for dot notation. -/ +@[reducible] def invertible.mul [monoid α] {a b : α} (ha : invertible a) (hb : invertible b) : + invertible (a * b) := +invertible_mul _ _ + theorem commute.inv_of_right [monoid α] {a b : α} [invertible b] (h : commute a b) : commute a (⅟b) := calc a * (⅟b) = (⅟b) * (b * a * (⅟b)) : by simp [mul_assoc] @@ -220,6 +232,43 @@ lemma nonzero_of_invertible [mul_zero_one_class α] (a : α) [nontrivial α] [in @[priority 100] instance invertible.ne_zero [mul_zero_one_class α] [nontrivial α] (a : α) [invertible a] : ne_zero a := ⟨nonzero_of_invertible a⟩ +section monoid +variables [monoid α] + +/-- This is the `invertible` version of `units.is_unit_units_mul` -/ +@[reducible] def invertible_of_invertible_mul (a b : α) [invertible a] [invertible (a * b)] : + invertible b := +{ inv_of := ⅟(a * b) * a, + inv_of_mul_self := by rw [mul_assoc, inv_of_mul_self], + mul_inv_of_self := by rw [←(is_unit_of_invertible a).mul_right_inj, ←mul_assoc, ←mul_assoc, + mul_inv_of_self, mul_one, one_mul] } + +/-- This is the `invertible` version of `units.is_unit_mul_units` -/ +@[reducible] def invertible_of_mul_invertible (a b : α) [invertible (a * b)] [invertible b] : + invertible a := +{ inv_of := b * ⅟(a * b), + inv_of_mul_self := by rw [←(is_unit_of_invertible b).mul_left_inj, mul_assoc, mul_assoc, + inv_of_mul_self, mul_one, one_mul], + mul_inv_of_self := by rw [←mul_assoc, mul_inv_of_self] } + +/-- `invertible_of_invertible_mul` and `invertible_mul` as an equivalence. -/ +@[simps] def invertible.mul_left {a : α} (ha : invertible a) (b : α) : + invertible b ≃ invertible (a * b) := +{ to_fun := λ hb, by exactI invertible_mul a b, + inv_fun := λ hab, by exactI invertible_of_invertible_mul a _, + left_inv := λ hb, subsingleton.elim _ _, + right_inv := λ hab, subsingleton.elim _ _, } + +/-- `invertible_of_mul_invertible` and `invertible_mul` as an equivalence. -/ +@[simps] def invertible.mul_right (a : α) {b : α} (ha : invertible b) : + invertible a ≃ invertible (a * b) := +{ to_fun := λ hb, by exactI invertible_mul a b, + inv_fun := λ hab, by exactI invertible_of_mul_invertible _ b, + left_inv := λ hb, subsingleton.elim _ _, + right_inv := λ hab, subsingleton.elim _ _, } + +end monoid + section monoid_with_zero variable [monoid_with_zero α] diff --git a/src/data/matrix/invertible.lean b/src/data/matrix/invertible.lean new file mode 100644 index 0000000000000..0ff393baefcf3 --- /dev/null +++ b/src/data/matrix/invertible.lean @@ -0,0 +1,91 @@ +/- +Copyright (c) 2023 Eric Wieser. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Eric Wieser +-/ +import algebra.invertible +import data.matrix.basic + +/-! # Extra lemmas about invertible matrices + +Many of the `invertible` lemmas are about `*`; this restates them to be about `⬝`. + +For lemmas about the matrix inverse in terms of the determinant and adjugate, see `matrix.has_inv` +in `linear_algebra/matrix/nonsingular_inverse.lean`. +-/ + +open_locale matrix + +variables {m n : Type*} {α : Type*} +variables [fintype n] [decidable_eq n] [semiring α] + +namespace matrix + +/-- A copy of `inv_of_mul_self` using `⬝` not `*`. -/ +protected lemma inv_of_mul_self (A : matrix n n α) [invertible A] : ⅟A ⬝ A = 1 := inv_of_mul_self A + +/-- A copy of `mul_inv_of_self` using `⬝` not `*`. -/ +protected lemma mul_inv_of_self (A : matrix n n α) [invertible A] : A ⬝ ⅟A = 1 := mul_inv_of_self A + +/-- A copy of `inv_of_mul_self_assoc` using `⬝` not `*`. -/ +protected lemma inv_of_mul_self_assoc (A : matrix n n α) (B : matrix n m α) [invertible A] : + ⅟A ⬝ (A ⬝ B) = B := +by rw [←matrix.mul_assoc, matrix.inv_of_mul_self, matrix.one_mul] + +/-- A copy of `mul_inv_of_self_assoc` using `⬝` not `*`. -/ +protected lemma mul_inv_of_self_assoc (A : matrix n n α) (B : matrix n m α) [invertible A] : + A ⬝ (⅟A ⬝ B) = B := +by rw [←matrix.mul_assoc, matrix.mul_inv_of_self, matrix.one_mul] + +/-- A copy of `mul_inv_of_mul_self_cancel` using `⬝` not `*`. -/ +protected lemma mul_inv_of_mul_self_cancel (A : matrix m n α) (B : matrix n n α) + [invertible B] : A ⬝ ⅟B ⬝ B = A := +by rw [matrix.mul_assoc, matrix.inv_of_mul_self, matrix.mul_one] + +/-- A copy of `mul_mul_inv_of_self_cancel` using `⬝` not `*`. -/ +protected lemma mul_mul_inv_of_self_cancel (A : matrix m n α) (B : matrix n n α) + [invertible B] : A ⬝ B ⬝ ⅟B = A := +by rw [matrix.mul_assoc, matrix.mul_inv_of_self, matrix.mul_one] + +/-- A copy of `invertible_mul` using `⬝` not `*`. -/ +@[reducible] protected def invertible_mul (A B : matrix n n α) [invertible A] [invertible B] : + invertible (A ⬝ B) := +{ inv_of := ⅟B ⬝ ⅟A, ..invertible_mul _ _ } + +/-- A copy of `invertible.mul` using `⬝` not `*`.-/ +@[reducible] def _root_.invertible.matrix_mul {A B : matrix n n α} + (ha : invertible A) (hb : invertible B) : invertible (A ⬝ B) := +invertible_mul _ _ + +protected lemma inv_of_mul {A B : matrix n n α} [invertible A] [invertible B] [invertible (A ⬝ B)] : + ⅟(A ⬝ B) = ⅟B ⬝ ⅟A := inv_of_mul _ _ + +/-- A copy of `invertible_of_invertible_mul` using `⬝` not `*`. -/ +@[reducible] protected def invertible_of_invertible_mul (a b : matrix n n α) + [invertible a] [invertible (a ⬝ b)] : invertible b := +{ inv_of := ⅟(a ⬝ b) ⬝ a, + ..invertible_of_invertible_mul a b } + +/-- A copy of `invertible_of_mul_invertible` using `⬝` not `*`. -/ +@[reducible] protected def invertible_of_mul_invertible (a b : matrix n n α) + [invertible (a ⬝ b)] [invertible b] : invertible a := +{ inv_of := b ⬝ ⅟(a ⬝ b), + ..invertible_of_mul_invertible a b } + +end matrix + +/-- A copy of `invertible.mul_left` using `⬝` not `*`. -/ +@[reducible] def invertible.matrix_mul_left + {a : matrix n n α} (ha : invertible a) (b : matrix n n α) : invertible b ≃ invertible (a ⬝ b) := +{ to_fun := λ hb, by exactI matrix.invertible_mul a b, + inv_fun := λ hab, by exactI matrix.invertible_of_invertible_mul a _, + left_inv := λ hb, subsingleton.elim _ _, + right_inv := λ hab, subsingleton.elim _ _, } + +/-- A copy of `invertible.mul_right` using `⬝` not `*`. -/ +@[reducible] def invertible.matrix_mul_right + (a : matrix n n α) {b : matrix n n α} (ha : invertible b) : invertible a ≃ invertible (a ⬝ b) := +{ to_fun := λ hb, by exactI matrix.invertible_mul a b, + inv_fun := λ hab, by exactI matrix.invertible_of_mul_invertible _ b, + left_inv := λ hb, subsingleton.elim _ _, + right_inv := λ hab, subsingleton.elim _ _, } diff --git a/src/linear_algebra/matrix/nonsingular_inverse.lean b/src/linear_algebra/matrix/nonsingular_inverse.lean index 66b7ddd6cdcaa..8a13e1a7fa31a 100644 --- a/src/linear_algebra/matrix/nonsingular_inverse.lean +++ b/src/linear_algebra/matrix/nonsingular_inverse.lean @@ -3,6 +3,7 @@ Copyright (c) 2019 Tim Baanen. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Tim Baanen, Lu-Ming Zhang -/ +import data.matrix.invertible import linear_algebra.matrix.adjugate /-! @@ -60,31 +61,6 @@ open equiv equiv.perm finset section invertible variables [fintype n] [decidable_eq n] [comm_ring α] -/-- A copy of `inv_of_mul_self` using `⬝` not `*`. -/ -protected lemma inv_of_mul_self (A : matrix n n α) [invertible A] : ⅟A ⬝ A = 1 := inv_of_mul_self A - -/-- A copy of `mul_inv_of_self` using `⬝` not `*`. -/ -protected lemma mul_inv_of_self (A : matrix n n α) [invertible A] : A ⬝ ⅟A = 1 := mul_inv_of_self A - -/-- A copy of `inv_of_mul_self_assoc` using `⬝` not `*`. -/ -protected lemma inv_of_mul_self_assoc (A : matrix n n α) (B : matrix n m α) [invertible A] : - ⅟A ⬝ (A ⬝ B) = B := -by rw [←matrix.mul_assoc, matrix.inv_of_mul_self, matrix.one_mul] - -/-- A copy of `mul_inv_of_self_assoc` using `⬝` not `*`. -/ -protected lemma mul_inv_of_self_assoc (A : matrix n n α) (B : matrix n m α) [invertible A] : - A ⬝ (⅟A ⬝ B) = B := -by rw [←matrix.mul_assoc, matrix.mul_inv_of_self, matrix.one_mul] - -/-- A copy of `mul_inv_of_mul_self_cancel` using `⬝` not `*`. -/ -protected lemma mul_inv_of_mul_self_cancel (A : matrix m n α) (B : matrix n n α) - [invertible B] : A ⬝ ⅟B ⬝ B = A := -by rw [matrix.mul_assoc, matrix.inv_of_mul_self, matrix.mul_one] - -/-- A copy of `mul_mul_inv_of_self_cancel` using `⬝` not `*`. -/ -protected lemma mul_mul_inv_of_self_cancel (A : matrix m n α) (B : matrix n n α) - [invertible B] : A ⬝ B ⬝ ⅟B = A := -by rw [matrix.mul_assoc, matrix.mul_inv_of_self, matrix.mul_one] variables (A : matrix n n α) (B : matrix n n α) From e8e130de9dba4ed6897183c3193c752ffadbcc77 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Wed, 21 Jun 2023 06:09:43 +0000 Subject: [PATCH 1206/1291] chore(*): add mathlib4 synchronization comments (#19208) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebraic_geometry.locally_ringed_space.has_colimits` * `analysis.bounded_variation` * `analysis.fourier.add_circle` * `analysis.fourier.poisson_summation` * `probability.borel_cantelli` * `probability.kernel.disintegration` * `probability.martingale.borel_cantelli` * `probability.martingale.convergence` * `probability.martingale.optional_stopping` * `probability.martingale.upcrossing` * `ring_theory.dedekind_domain.integral_closure` * `ring_theory.localization.norm` * `ring_theory.norm` * `ring_theory.polynomial.eisenstein.is_integral` * `ring_theory.trace` --- src/algebraic_geometry/locally_ringed_space/has_colimits.lean | 3 +++ src/analysis/bounded_variation.lean | 3 +++ src/analysis/fourier/add_circle.lean | 3 +++ src/analysis/fourier/poisson_summation.lean | 3 +++ src/probability/borel_cantelli.lean | 3 +++ src/probability/kernel/disintegration.lean | 3 +++ src/probability/martingale/borel_cantelli.lean | 3 +++ src/probability/martingale/convergence.lean | 3 +++ src/probability/martingale/optional_stopping.lean | 3 +++ src/probability/martingale/upcrossing.lean | 3 +++ src/ring_theory/dedekind_domain/integral_closure.lean | 3 +++ src/ring_theory/localization/norm.lean | 3 +++ src/ring_theory/norm.lean | 3 +++ src/ring_theory/polynomial/eisenstein/is_integral.lean | 3 +++ src/ring_theory/trace.lean | 3 +++ 15 files changed, 45 insertions(+) diff --git a/src/algebraic_geometry/locally_ringed_space/has_colimits.lean b/src/algebraic_geometry/locally_ringed_space/has_colimits.lean index 51c20a36cb793..a25dffcdd9fc3 100644 --- a/src/algebraic_geometry/locally_ringed_space/has_colimits.lean +++ b/src/algebraic_geometry/locally_ringed_space/has_colimits.lean @@ -11,6 +11,9 @@ import category_theory.limits.constructions.limits_of_products_and_equalizers /-! # Colimits of LocallyRingedSpace +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We construct the explicit coproducts and coequalizers of `LocallyRingedSpace`. It then follows that `LocallyRingedSpace` has all colimits, and `forget_to_SheafedSpace` preserves them. diff --git a/src/analysis/bounded_variation.lean b/src/analysis/bounded_variation.lean index 1fae789f4cb82..252ed7aa1747e 100644 --- a/src/analysis/bounded_variation.lean +++ b/src/analysis/bounded_variation.lean @@ -14,6 +14,9 @@ import tactic.wlog /-! # Functions of bounded variation +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We study functions of bounded variation. In particular, we show that a bounded variation function is a difference of monotone functions, and differentiable almost everywhere. This implies that Lipschitz functions from the real line into finite-dimensional vector space are also differentiable diff --git a/src/analysis/fourier/add_circle.lean b/src/analysis/fourier/add_circle.lean index 7b754f1353ec4..a5e5f3877cfdd 100644 --- a/src/analysis/fourier/add_circle.lean +++ b/src/analysis/fourier/add_circle.lean @@ -17,6 +17,9 @@ import measure_theory.integral.fund_thm_calculus # Fourier analysis on the additive circle +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains basic results on Fourier series for functions on the additive circle `add_circle T = ℝ / ℤ • T`. diff --git a/src/analysis/fourier/poisson_summation.lean b/src/analysis/fourier/poisson_summation.lean index 04cff9f36353d..7c7f257793639 100644 --- a/src/analysis/fourier/poisson_summation.lean +++ b/src/analysis/fourier/poisson_summation.lean @@ -12,6 +12,9 @@ import measure_theory.measure.lebesgue.integral /-! # Poisson's summation formula +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We prove Poisson's summation formula `∑ (n : ℤ), f n = ∑ (n : ℤ), 𝓕 f n`, where `𝓕 f` is the Fourier transform of `f`, under the following hypotheses: * `f` is a continuous function `ℝ → ℂ`. diff --git a/src/probability/borel_cantelli.lean b/src/probability/borel_cantelli.lean index 466f4da94528b..47b794c3ddc16 100644 --- a/src/probability/borel_cantelli.lean +++ b/src/probability/borel_cantelli.lean @@ -11,6 +11,9 @@ import probability.independence.basic # The second Borel-Cantelli lemma +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains the second Borel-Cantelli lemma which states that, given a sequence of independent sets `(sₙ)` in a probability space, if `∑ n, μ sₙ = ∞`, then the limsup of `sₙ` has measure 1. We employ a proof using Lévy's generalized Borel-Cantelli by choosing an appropriate diff --git a/src/probability/kernel/disintegration.lean b/src/probability/kernel/disintegration.lean index a4cf308924105..3636f104864f4 100644 --- a/src/probability/kernel/disintegration.lean +++ b/src/probability/kernel/disintegration.lean @@ -10,6 +10,9 @@ import probability.kernel.integral_comp_prod /-! # Disintegration of measures on product spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Let `ρ` be a finite measure on `α × Ω`, where `Ω` is a standard Borel space. In mathlib terms, `Ω` verifies `[nonempty Ω] [topological_space Ω] [polish_space Ω] [measurable_space Ω] [borel_space Ω]`. Then there exists a kernel `ρ.cond_kernel : kernel α Ω` such that for any measurable set diff --git a/src/probability/martingale/borel_cantelli.lean b/src/probability/martingale/borel_cantelli.lean index fb9014347322b..4a97b49a346b9 100644 --- a/src/probability/martingale/borel_cantelli.lean +++ b/src/probability/martingale/borel_cantelli.lean @@ -11,6 +11,9 @@ import probability.martingale.centering # Generalized Borel-Cantelli lemma +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves Lévy's generalized Borel-Cantelli lemma which is a generalization of the Borel-Cantelli lemmas. With this generalization, one can easily deduce the Borel-Cantelli lemmas by choosing appropriate filtrations. This file also contains the one sided martingale bound which diff --git a/src/probability/martingale/convergence.lean b/src/probability/martingale/convergence.lean index f670120fad772..340f3fba23564 100644 --- a/src/probability/martingale/convergence.lean +++ b/src/probability/martingale/convergence.lean @@ -11,6 +11,9 @@ import measure_theory.constructions.polish # Martingale convergence theorems +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The martingale convergence theorems are a collection of theorems characterizing the convergence of a martingale provided it satisfies some boundedness conditions. This file contains the almost everywhere martingale convergence theorem which provides an almost everywhere limit to diff --git a/src/probability/martingale/optional_stopping.lean b/src/probability/martingale/optional_stopping.lean index 5e609db29e41e..3645319aa2ac2 100644 --- a/src/probability/martingale/optional_stopping.lean +++ b/src/probability/martingale/optional_stopping.lean @@ -8,6 +8,9 @@ import probability.martingale.basic /-! # Optional stopping theorem (fair game theorem) +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The optional stopping theorem states that an adapted integrable process `f` is a submartingale if and only if for all bounded stopping times `τ` and `π` such that `τ ≤ π`, the stopped value of `f` at `τ` has expectation smaller than its stopped value at `π`. diff --git a/src/probability/martingale/upcrossing.lean b/src/probability/martingale/upcrossing.lean index 2bc98deb4de6a..8f0acc6e431a7 100644 --- a/src/probability/martingale/upcrossing.lean +++ b/src/probability/martingale/upcrossing.lean @@ -11,6 +11,9 @@ import probability.martingale.basic # Doob's upcrossing estimate +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given a discrete real-valued submartingale $(f_n)_{n \in \mathbb{N}}$, denoting $U_N(a, b)$ the number of times $f_n$ crossed from below $a$ to above $b$ before time $N$, Doob's upcrossing estimate (also known as Doob's inequality) states that diff --git a/src/ring_theory/dedekind_domain/integral_closure.lean b/src/ring_theory/dedekind_domain/integral_closure.lean index 8116312cb9f07..357d4b6ec65e0 100644 --- a/src/ring_theory/dedekind_domain/integral_closure.lean +++ b/src/ring_theory/dedekind_domain/integral_closure.lean @@ -11,6 +11,9 @@ import ring_theory.trace /-! # Integral closure of Dedekind domains +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file shows the integral closure of a Dedekind domain (in particular, the ring of integers of a number field) is a Dedekind domain. diff --git a/src/ring_theory/localization/norm.lean b/src/ring_theory/localization/norm.lean index 5eba4dbb697ea..641c9a51d323d 100644 --- a/src/ring_theory/localization/norm.lean +++ b/src/ring_theory/localization/norm.lean @@ -11,6 +11,9 @@ import ring_theory.norm # Field/algebra norm and localization +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains results on the combination of `algebra.norm` and `is_localization`. ## Main results diff --git a/src/ring_theory/norm.lean b/src/ring_theory/norm.lean index d8637f103a938..5a8687eeb9bad 100644 --- a/src/ring_theory/norm.lean +++ b/src/ring_theory/norm.lean @@ -15,6 +15,9 @@ import field_theory.galois /-! # Norm for (finite) ring extensions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Suppose we have an `R`-algebra `S` with a finite basis. For each `s : S`, the determinant of the linear map given by multiplying by `s` gives information about the roots of the minimal polynomial of `s` over `R`. diff --git a/src/ring_theory/polynomial/eisenstein/is_integral.lean b/src/ring_theory/polynomial/eisenstein/is_integral.lean index 0704e82c63a77..ca3651eddda98 100644 --- a/src/ring_theory/polynomial/eisenstein/is_integral.lean +++ b/src/ring_theory/polynomial/eisenstein/is_integral.lean @@ -10,6 +10,9 @@ import ring_theory.polynomial.cyclotomic.expand /-! # Eisenstein polynomials + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. In this file we gather more miscellaneous results about Eisenstein polynomials ## Main results diff --git a/src/ring_theory/trace.lean b/src/ring_theory/trace.lean index 5f16361c7da38..c5d0ce8bf0fdf 100644 --- a/src/ring_theory/trace.lean +++ b/src/ring_theory/trace.lean @@ -18,6 +18,9 @@ import ring_theory.power_basis /-! # Trace for (finite) ring extensions. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Suppose we have an `R`-algebra `S` with a finite basis. For each `s : S`, the trace of the linear map given by multiplying by `s` gives information about the roots of the minimal polynomial of `s` over `R`. From c14c8fcde993801fca8946b0d80131a1a81d1520 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 21 Jun 2023 07:24:58 +0000 Subject: [PATCH 1207/1291] feat(measure_theory/integral/average): Lebesgue average (#19199) Define the Lebesgue integral version of the average of a measurable function and prove the corresponding first moment method. --- src/data/real/ennreal.lean | 13 + src/measure_theory/integral/average.lean | 353 ++++++++++++++++-- src/measure_theory/integral/lebesgue.lean | 24 +- src/measure_theory/measurable_space.lean | 2 +- .../measure/measure_space_def.lean | 6 +- src/probability/density.lean | 2 +- 6 files changed, 369 insertions(+), 31 deletions(-) diff --git a/src/data/real/ennreal.lean b/src/data/real/ennreal.lean index 00d67e43847e3..b01dfe825be35 100644 --- a/src/data/real/ennreal.lean +++ b/src/data/real/ennreal.lean @@ -105,6 +105,10 @@ noncomputable instance : linear_ordered_comm_monoid_with_zero ℝ≥0∞ := .. ennreal.linear_ordered_add_comm_monoid_with_top, .. (show comm_semiring ℝ≥0∞, from infer_instance) } +instance : unique (add_units ℝ≥0∞) := +{ default := 0, + uniq := λ a, add_units.ext $ le_zero_iff.1 $ by { rw ←a.add_neg, exact le_self_add } } + instance : inhabited ℝ≥0∞ := ⟨0⟩ instance : has_coe ℝ≥0 ℝ≥0∞ := ⟨ option.some ⟩ @@ -1056,6 +1060,12 @@ by rw [div_eq_mul_inv, mul_assoc, ennreal.inv_mul_cancel h0 hI, mul_one] protected lemma mul_div_cancel' (h0 : a ≠ 0) (hI : a ≠ ∞) : a * (b / a) = b := by rw [mul_comm, ennreal.div_mul_cancel h0 hI] +protected lemma mul_comm_div : a / b * c = a * (c / b) := +by simp only [div_eq_mul_inv, mul_comm, mul_assoc] + +protected lemma mul_div_right_comm : a * b / c = a / c * b := +by simp only [div_eq_mul_inv, mul_comm, mul_assoc] + instance : has_involutive_inv ℝ≥0∞ := { inv := has_inv.inv, inv_inv := λ a, by @@ -1077,6 +1087,9 @@ inv_top ▸ inv_inj protected lemma inv_ne_zero : a⁻¹ ≠ 0 ↔ a ≠ ∞ := by simp +protected lemma div_pos (ha : a ≠ 0) (hb : b ≠ ∞) : 0 < a / b := +ennreal.mul_pos ha $ ennreal.inv_ne_zero.2 hb + protected lemma mul_inv {a b : ℝ≥0∞} (ha : a ≠ 0 ∨ b ≠ ∞) (hb : a ≠ ∞ ∨ b ≠ 0) : (a * b)⁻¹ = a⁻¹ * b⁻¹ := begin diff --git a/src/measure_theory/integral/average.lean b/src/measure_theory/integral/average.lean index 230993042e849..05d85570cef2b 100644 --- a/src/measure_theory/integral/average.lean +++ b/src/measure_theory/integral/average.lean @@ -19,6 +19,8 @@ measure, then the average of any function is equal to its integral. For the average on a set, we use `⨍ x in s, f x ∂μ` (notation for `⨍ x, f x ∂(μ.restrict s)`). For average w.r.t. the volume, one can omit `∂volume`. +Both have a version for the Lebesgue integral rather than Bochner. + We prove several version of the first moment method: An integrable function is below/above its average on a set of positive measure. @@ -28,11 +30,6 @@ The average is defined as an integral over `(μ univ)⁻¹ • μ` so that all t integrals work for the average without modifications. For theorems that require integrability of a function, we provide a convenience lemma `measure_theory.integrable.to_average`. -## TODO - -Provide the first moment method for the Lebesgue integral as well. A draft is available on branch -`first_moment_lintegral`. - ## Tags integral, center mass, average value @@ -44,19 +41,174 @@ open_locale topology big_operators ennreal convex variables {α E F : Type*} {m0 : measurable_space α} [normed_add_comm_group E] [normed_space ℝ E] [complete_space E] [normed_add_comm_group F] [normed_space ℝ F] [complete_space F] - {μ : measure α} {s t : set α} + {μ ν : measure α} {s t : set α} /-! ### Average value of a function w.r.t. a measure -The average value of a function `f` w.r.t. a measure `μ` (notation: `⨍ x, f x ∂μ`) is defined as -`(μ univ).to_real⁻¹ • ∫ x, f x ∂μ`, so it is equal to zero if `f` is not integrable or if `μ` is an -infinite measure. If `μ` is a probability measure, then the average of any function is equal to its -integral. - +The (Bochner, Lebesgue) average value of a function `f` w.r.t. a measure `μ` (notation: +`⨍ x, f x ∂μ`, `⨍⁻ x, f x ∂μ`) is defined as the (Bochner, Lebesgue) integral divided by the total +measure, so it is equal to zero if `μ` is an infinite measure, and (typically) equal to infinity if +`f` is not integrable. If `μ` is a probability measure, then the average of any function is equal to +its integral. -/ namespace measure_theory +section ennreal +variables (μ) {f g : α → ℝ≥0∞} +include m0 + +/-- Average value of an `ℝ≥0∞`-valued function `f` w.r.t. a measure `μ`, notation: `⨍⁻ x, f x ∂μ`. +It is defined as `μ univ⁻¹ * ∫⁻ x, f x ∂μ`, so it is equal to zero if `μ` is an infinite measure. If +`μ` is a probability measure, then the average of any function is equal to its integral. + +For the average on a set, use `⨍⁻ x in s, f x ∂μ` (defined as `⨍⁻ x, f x ∂(μ.restrict s)`). For +average w.r.t. the volume, one can omit `∂volume`. -/ +noncomputable def laverage (f : α → ℝ≥0∞) := ∫⁻ x, f x ∂((μ univ)⁻¹ • μ) + +notation `⨍⁻` binders `, ` r:(scoped:60 f, f) ` ∂` μ:70 := laverage μ r +notation `⨍⁻` binders `, ` r:(scoped:60 f, laverage volume f) := r +notation `⨍⁻` binders ` in ` s `, ` r:(scoped:60 f, f) ` ∂` μ:70 := + laverage (measure.restrict μ s) r +notation `⨍⁻` binders ` in ` s `, ` r:(scoped:60 f, laverage (measure.restrict volume s) f) := r + +@[simp] lemma laverage_zero : ⨍⁻ x, (0 : ℝ≥0∞) ∂μ = 0 := by rw [laverage, lintegral_zero] + +@[simp] lemma laverage_zero_measure (f : α → ℝ≥0∞) : ⨍⁻ x, f x ∂(0 : measure α) = 0 := +by simp [laverage] + +lemma laverage_eq' (f : α → ℝ≥0∞) : ⨍⁻ x, f x ∂μ = ∫⁻ x, f x ∂((μ univ)⁻¹ • μ) := rfl + +lemma laverage_eq (f : α → ℝ≥0∞) : ⨍⁻ x, f x ∂μ = ∫⁻ x, f x ∂μ / μ univ := +by rw [laverage_eq', lintegral_smul_measure, ennreal.div_eq_inv_mul] + +lemma laverage_eq_lintegral [is_probability_measure μ] (f : α → ℝ≥0∞) : + ⨍⁻ x, f x ∂μ = ∫⁻ x, f x ∂μ := +by rw [laverage, measure_univ, inv_one, one_smul] + +@[simp] lemma measure_mul_laverage [is_finite_measure μ] (f : α → ℝ≥0∞) : + μ univ * ⨍⁻ x, f x ∂μ = ∫⁻ x, f x ∂μ := +begin + cases eq_or_ne μ 0 with hμ hμ, + { rw [hμ, lintegral_zero_measure, laverage_zero_measure, mul_zero] }, + { rw [laverage_eq, ennreal.mul_div_cancel' (measure_univ_ne_zero.2 hμ) (measure_ne_top _ _)] } +end + +lemma set_laverage_eq (f : α → ℝ≥0∞) (s : set α) : ⨍⁻ x in s, f x ∂μ = ∫⁻ x in s, f x ∂μ / μ s := +by rw [laverage_eq, restrict_apply_univ] + +lemma set_laverage_eq' (f : α → ℝ≥0∞) (s : set α) : + ⨍⁻ x in s, f x ∂μ = ∫⁻ x, f x ∂((μ s)⁻¹ • μ.restrict s) := +by simp only [laverage_eq', restrict_apply_univ] + +variable {μ} + +lemma laverage_congr {f g : α → ℝ≥0∞} (h : f =ᵐ[μ] g) : ⨍⁻ x, f x ∂μ = ⨍⁻ x, g x ∂μ := +by simp only [laverage_eq, lintegral_congr_ae h] + +lemma set_laverage_congr (h : s =ᵐ[μ] t) : ⨍⁻ x in s, f x ∂μ = ⨍⁻ x in t, f x ∂μ := +by simp only [set_laverage_eq, set_lintegral_congr h, measure_congr h] + +lemma set_laverage_congr_fun (hs : measurable_set s) (h : ∀ᵐ x ∂μ, x ∈ s → f x = g x) : + ⨍⁻ x in s, f x ∂μ = ⨍⁻ x in s, g x ∂μ := +by simp only [laverage_eq, set_lintegral_congr_fun hs h] + +lemma laverage_lt_top (hf : ∫⁻ x, f x ∂μ ≠ ∞) : ⨍⁻ x, f x ∂μ < ∞ := +begin + obtain rfl | hμ := eq_or_ne μ 0, + { simp }, + { rw laverage_eq, + exact div_lt_top hf (measure_univ_ne_zero.2 hμ) } +end + +lemma set_laverage_lt_top : ∫⁻ x in s, f x ∂μ ≠ ∞ → ⨍⁻ x in s, f x ∂μ < ∞ := laverage_lt_top + +lemma laverage_add_measure : + ⨍⁻ x, f x ∂(μ + ν) = + μ univ / (μ univ + ν univ) * ⨍⁻ x, f x ∂μ + ν univ / (μ univ + ν univ) * ⨍⁻ x, f x ∂ν := +begin + by_cases hμ : is_finite_measure μ, swap, + { rw not_is_finite_measure_iff at hμ, + simp [laverage_eq, hμ] }, + by_cases hν : is_finite_measure ν, swap, + { rw not_is_finite_measure_iff at hν, + simp [laverage_eq, hν] }, + haveI := hμ, haveI := hν, + simp only [←ennreal.mul_div_right_comm, measure_mul_laverage, ←ennreal.add_div, + ←lintegral_add_measure, ←measure.add_apply, ←laverage_eq], +end + +lemma measure_mul_set_laverage (f : α → ℝ≥0∞) (h : μ s ≠ ∞) : + μ s * ⨍⁻ x in s, f x ∂μ = ∫⁻ x in s, f x ∂μ := +by { haveI := fact.mk h.lt_top, rw [← measure_mul_laverage, restrict_apply_univ] } + +lemma laverage_union (hd : ae_disjoint μ s t) (ht : null_measurable_set t μ) : + ⨍⁻ x in s ∪ t, f x ∂μ = + μ s / (μ s + μ t) * ⨍⁻ x in s, f x ∂μ + μ t / (μ s + μ t) * ⨍⁻ x in t, f x ∂μ := +by rw [restrict_union₀ hd ht, laverage_add_measure, restrict_apply_univ, restrict_apply_univ] + +lemma laverage_union_mem_open_segment (hd : ae_disjoint μ s t) (ht : null_measurable_set t μ) + (hs₀ : μ s ≠ 0) (ht₀ : μ t ≠ 0) (hsμ : μ s ≠ ∞) (htμ : μ t ≠ ∞) : + ⨍⁻ x in s ∪ t, f x ∂μ ∈ open_segment ℝ≥0∞ (⨍⁻ x in s, f x ∂μ) (⨍⁻ x in t, f x ∂μ) := +begin + refine ⟨μ s / (μ s + μ t), μ t / (μ s + μ t), ennreal.div_pos hs₀ $ add_ne_top.2 ⟨hsμ, htμ⟩, + ennreal.div_pos ht₀ $ add_ne_top.2 ⟨hsμ, htμ⟩, _, (laverage_union hd ht).symm⟩, + rw [←ennreal.add_div, ennreal.div_self (add_eq_zero.not.2 $ λ h, hs₀ h.1) + (add_ne_top.2 ⟨hsμ, htμ⟩)], +end + +lemma laverage_union_mem_segment (hd : ae_disjoint μ s t) (ht : null_measurable_set t μ) + (hsμ : μ s ≠ ∞) (htμ : μ t ≠ ∞) : + ⨍⁻ x in s ∪ t, f x ∂μ ∈ [⨍⁻ x in s, f x ∂μ -[ℝ≥0∞] ⨍⁻ x in t, f x ∂μ] := +begin + by_cases hs₀ : μ s = 0, + { rw ← ae_eq_empty at hs₀, + rw [restrict_congr_set (hs₀.union eventually_eq.rfl), empty_union], + exact right_mem_segment _ _ _ }, + { refine ⟨μ s / (μ s + μ t), μ t / (μ s + μ t), zero_le _, zero_le _, _, + (laverage_union hd ht).symm⟩, + rw [←ennreal.add_div, ennreal.div_self (add_eq_zero.not.2 $ λ h, hs₀ h.1) + (add_ne_top.2 ⟨hsμ, htμ⟩)] } +end + +lemma laverage_mem_open_segment_compl_self [is_finite_measure μ] (hs : null_measurable_set s μ) + (hs₀ : μ s ≠ 0) (hsc₀ : μ sᶜ ≠ 0) : + ⨍⁻ x, f x ∂μ ∈ open_segment ℝ≥0∞ (⨍⁻ x in s, f x ∂μ) (⨍⁻ x in sᶜ, f x ∂μ) := +by simpa only [union_compl_self, restrict_univ] + using laverage_union_mem_open_segment ae_disjoint_compl_right hs.compl hs₀ hsc₀ + (measure_ne_top _ _) (measure_ne_top _ _) + +@[simp] lemma laverage_const (μ : measure α) [is_finite_measure μ] [h : μ.ae.ne_bot] (c : ℝ≥0∞) : + ⨍⁻ x, c ∂μ = c := +by simp only [laverage_eq, lintegral_const, measure.restrict_apply, measurable_set.univ, univ_inter, + div_eq_mul_inv, mul_assoc, ennreal.mul_inv_cancel, mul_one, measure_ne_top μ univ, ne.def, + measure_univ_ne_zero, ae_ne_bot.1 h, not_false_iff] + +lemma set_laverage_const (hs₀ : μ s ≠ 0) (hs : μ s ≠ ∞) (c : ℝ≥0∞) : ⨍⁻ x in s, c ∂μ = c := +by simp only [set_laverage_eq, lintegral_const, measure.restrict_apply, measurable_set.univ, + univ_inter, div_eq_mul_inv, mul_assoc, ennreal.mul_inv_cancel hs₀ hs, mul_one] + +@[simp] lemma laverage_one [is_finite_measure μ] [h : μ.ae.ne_bot] : ⨍⁻ x, (1 : ℝ≥0∞) ∂μ = 1 := +laverage_const _ _ + +lemma set_laverage_one (hs₀ : μ s ≠ 0) (hs : μ s ≠ ∞) : ⨍⁻ x in s, (1 : ℝ≥0∞) ∂μ = 1 := +set_laverage_const hs₀ hs _ + +@[simp] lemma lintegral_laverage (μ : measure α) [is_finite_measure μ] (f : α → ℝ≥0∞) : + ∫⁻ x, ⨍⁻ a, f a ∂μ ∂μ = ∫⁻ x, f x ∂μ := +begin + unfreezingI { obtain rfl | hμ := eq_or_ne μ 0 }, + { simp }, + { rw [lintegral_const, laverage_eq, + ennreal.div_mul_cancel (measure_univ_ne_zero.2 hμ) (measure_ne_top _ _)] } +end + +lemma set_lintegral_set_laverage (μ : measure α) [is_finite_measure μ] (f : α → ℝ≥0∞) (s : set α) : + ∫⁻ x in s, ⨍⁻ a in s, f a ∂μ ∂μ = ∫⁻ x in s, f x ∂μ := +lintegral_laverage _ _ + +end ennreal + section normed_add_comm_group variables (μ) {f g : α → E} include m0 @@ -114,9 +266,13 @@ variable {μ} lemma average_congr {f g : α → E} (h : f =ᵐ[μ] g) : ⨍ x, f x ∂μ = ⨍ x, g x ∂μ := by simp only [average_eq, integral_congr_ae h] -lemma set_average_congr_set_ae (h : s =ᵐ[μ] t) : ⨍ x in s, f x ∂μ = ⨍ x in t, f x ∂μ := +lemma set_average_congr (h : s =ᵐ[μ] t) : ⨍ x in s, f x ∂μ = ⨍ x in t, f x ∂μ := by simp only [set_average_eq, set_integral_congr_set_ae h, measure_congr h] +lemma set_average_congr_fun (hs : measurable_set s) (h : ∀ᵐ x ∂μ, x ∈ s → f x = g x) : + ⨍ x in s, f x ∂μ = ⨍ x in s, g x ∂μ := +by simp only [average_eq, set_integral_congr_ae hs h] + lemma average_add_measure [is_finite_measure μ] {ν : measure α} [is_finite_measure ν] {f : α → E} (hμ : integrable f μ) (hν : integrable f ν) : ⨍ x, f x ∂(μ + ν) = @@ -180,7 +336,7 @@ by simpa only [union_compl_self, restrict_univ] using average_union_mem_open_segment ae_disjoint_compl_right hs.compl hs₀ hsc₀ (measure_ne_top _ _) (measure_ne_top _ _) hfi.integrable_on hfi.integrable_on -@[simp] lemma average_const [is_finite_measure μ] [h : μ.ae.ne_bot] (c : E) : +@[simp] lemma average_const (μ : measure α) [is_finite_measure μ] [h : μ.ae.ne_bot] (c : E) : ⨍ x, c ∂μ = c := by simp only [average_eq, integral_const, measure.restrict_apply, measurable_set.univ, one_smul, univ_inter, smul_smul, ← ennreal.to_real_inv, ← ennreal.to_real_mul, ennreal.inv_mul_cancel, @@ -245,23 +401,23 @@ lemma of_real_set_average {f : α → ℝ} (hf : integrable_on f s μ) ennreal.of_real (⨍ x in s, f x ∂μ) = (∫⁻ x in s, ennreal.of_real (f x) ∂μ) / μ s := by simpa using of_real_average hf hf₀ -lemma average_to_real {f : α → ℝ≥0∞} (hf : ae_measurable f μ) (hf' : ∀ᵐ x ∂μ, f x ≠ ∞) : - ⨍ x, (f x).to_real ∂μ = (∫⁻ x, f x ∂μ / μ univ).to_real := +lemma to_real_laverage {f : α → ℝ≥0∞} (hf : ae_measurable f μ) (hf' : ∀ᵐ x ∂μ, f x ≠ ∞) : + (⨍⁻ x, f x ∂μ).to_real = ⨍ x, (f x).to_real ∂μ := begin obtain rfl | hμ := eq_or_ne μ 0, { simp }, - { rw [average_eq, smul_eq_mul, to_real_div, - ←integral_to_real hf (hf'.mono $ λ _, lt_top_iff_ne_top.2), div_eq_inv_mul] } + { rw [average_eq, laverage_eq, smul_eq_mul, to_real_div, div_eq_inv_mul, + ←integral_to_real hf (hf'.mono $ λ _, lt_top_iff_ne_top.2)] } end -lemma set_average_to_real {f : α → ℝ≥0∞} (hf : ae_measurable f (μ.restrict s)) +lemma to_real_set_laverage {f : α → ℝ≥0∞} (hf : ae_measurable f (μ.restrict s)) (hf' : ∀ᵐ x ∂(μ.restrict s), f x ≠ ∞) : - ⨍ x in s, (f x).to_real ∂μ = (∫⁻ x in s, f x ∂μ / μ s).to_real := -by simpa using average_to_real hf hf' + (∫⁻ x in s, f x ∂μ / μ s).to_real = ⨍ x in s, (f x).to_real ∂μ := +by simpa [laverage_eq] using to_real_laverage hf hf' /-! ### First moment method -/ -section first_moment +section first_moment_real variables {N : set α} {f : α → ℝ} /-- **First moment method**. An integrable function is smaller than its mean on a set of positive @@ -384,5 +540,158 @@ by simpa only [average_eq_integral] using exists_not_mem_null_average_le (is_probability_measure.ne_zero μ) hf hN end probability_measure -end first_moment +end first_moment_real + +section first_moment_ennreal +variables {N : set α} {f : α → ℝ≥0∞} + +/-- **First moment method**. A measurable function is smaller than its mean on a set of positive +measure. -/ +lemma measure_le_set_laverage_pos (hμ : μ s ≠ 0) (hμ₁ : μ s ≠ ∞) + (hf : ae_measurable f (μ.restrict s)) : 0 < μ {x ∈ s | f x ≤ ⨍⁻ a in s, f a ∂μ} := +begin + obtain h | h := eq_or_ne (∫⁻ a in s, f a ∂μ) ∞, + { simpa [mul_top, hμ₁, laverage, h, top_div_of_ne_top hμ₁, pos_iff_ne_zero] using hμ }, + have := measure_le_set_average_pos hμ hμ₁ (integrable_to_real_of_lintegral_ne_top hf h), + rw [←set_of_inter_eq_sep, ←measure.restrict_apply₀ + (hf.ae_strongly_measurable.null_measurable_set_le ae_strongly_measurable_const)], + rw [←set_of_inter_eq_sep, ←measure.restrict_apply₀ + (hf.ennreal_to_real.ae_strongly_measurable.null_measurable_set_le ae_strongly_measurable_const), + ←measure_diff_null (measure_eq_top_of_lintegral_ne_top hf h)] at this, + refine this.trans_le (measure_mono _), + rintro x ⟨hfx, hx⟩, + dsimp at hfx, + rwa [←to_real_laverage hf, to_real_le_to_real hx (set_laverage_lt_top h).ne] at hfx, + { simp_rw [ae_iff, not_ne_iff], + exact measure_eq_top_of_lintegral_ne_top hf h } +end + +/-- **First moment method**. A measurable function is greater than its mean on a set of positive +measure. -/ +lemma measure_set_laverage_le_pos (hμ : μ s ≠ 0) (hs : null_measurable_set s μ) + (hint : ∫⁻ a in s, f a ∂μ ≠ ∞) : 0 < μ {x ∈ s | ⨍⁻ a in s, f a ∂μ ≤ f x} := +begin + obtain hμ₁ | hμ₁ := eq_or_ne (μ s) ∞, + { simp [set_laverage_eq, hμ₁] }, + obtain ⟨g, hg, hgf, hfg⟩ := exists_measurable_le_lintegral_eq (μ.restrict s) f, + have hfg' : ⨍⁻ a in s, f a ∂μ = ⨍⁻ a in s, g a ∂μ, + { simp_rw [laverage_eq, hfg] }, + rw hfg at hint, + have := measure_set_average_le_pos hμ hμ₁ + (integrable_to_real_of_lintegral_ne_top hg.ae_measurable hint), + simp_rw [←set_of_inter_eq_sep, ←measure.restrict_apply₀' hs, hfg'], + rw [←set_of_inter_eq_sep, ←measure.restrict_apply₀' hs, + ←measure_diff_null (measure_eq_top_of_lintegral_ne_top hg.ae_measurable hint)] at this, + refine this.trans_le (measure_mono _), + rintro x ⟨hfx, hx⟩, + dsimp at hfx, + rw [←to_real_laverage hg.ae_measurable, to_real_le_to_real (set_laverage_lt_top hint).ne hx] + at hfx, + exact hfx.trans (hgf _), + { simp_rw [ae_iff, not_ne_iff], + exact measure_eq_top_of_lintegral_ne_top hg.ae_measurable hint } +end + +/-- **First moment method**. The minimum of a measurable function is smaller than its mean. -/ +lemma exists_le_set_laverage (hμ : μ s ≠ 0) (hμ₁ : μ s ≠ ∞) (hf : ae_measurable f (μ.restrict s)) : + ∃ x ∈ s, f x ≤ ⨍⁻ a in s, f a ∂μ := +let ⟨x, hx, h⟩ := nonempty_of_measure_ne_zero (measure_le_set_laverage_pos hμ hμ₁ hf).ne' + in ⟨x, hx, h⟩ + +/-- **First moment method**. The maximum of a measurable function is greater than its mean. -/ +lemma exists_set_laverage_le (hμ : μ s ≠ 0) (hs : null_measurable_set s μ) + (hint : ∫⁻ a in s, f a ∂μ ≠ ∞) : ∃ x ∈ s, ⨍⁻ a in s, f a ∂μ ≤ f x := +let ⟨x, hx, h⟩ := nonempty_of_measure_ne_zero (measure_set_laverage_le_pos hμ hs hint).ne' + in ⟨x, hx, h⟩ + +/-- **First moment method**. A measurable function is greater than its mean on a set of positive +measure. -/ +lemma measure_laverage_le_pos (hμ : μ ≠ 0) (hint : ∫⁻ a, f a ∂μ ≠ ∞) : + 0 < μ {x | ⨍⁻ a, f a ∂μ ≤ f x} := +by simpa [hint] using @measure_set_laverage_le_pos _ _ _ _ f (measure_univ_ne_zero.2 hμ) + null_measurable_set_univ + +/-- **First moment method**. The maximum of a measurable function is greater than its mean. -/ +lemma exists_laverage_le (hμ : μ ≠ 0) (hint : ∫⁻ a, f a ∂μ ≠ ∞) : ∃ x, ⨍⁻ a, f a ∂μ ≤ f x := +let ⟨x, hx⟩ := nonempty_of_measure_ne_zero (measure_laverage_le_pos hμ hint).ne' in ⟨x, hx⟩ + +/-- **First moment method**. The maximum of a measurable function is greater than its mean, while +avoiding a null set. -/ +lemma exists_not_mem_null_laverage_le (hμ : μ ≠ 0) (hint : ∫⁻ (a : α), f a ∂μ ≠ ∞) (hN : μ N = 0) : + ∃ x ∉ N, ⨍⁻ a, f a ∂μ ≤ f x := +begin + have := measure_laverage_le_pos hμ hint, + rw ←measure_diff_null hN at this, + obtain ⟨x, hx, hxN⟩ := nonempty_of_measure_ne_zero this.ne', + exact ⟨x, hxN, hx⟩, +end + +section finite_measure +variables [is_finite_measure μ] + +/-- **First moment method**. A measurable function is smaller than its mean on a set of positive +measure. -/ +lemma measure_le_laverage_pos (hμ : μ ≠ 0) (hf : ae_measurable f μ) : + 0 < μ {x | f x ≤ ⨍⁻ a, f a ∂μ} := +by simpa + using measure_le_set_laverage_pos (measure_univ_ne_zero.2 hμ) (measure_ne_top _ _) hf.restrict + +/-- **First moment method**. The minimum of a measurable function is smaller than its mean. -/ +lemma exists_le_laverage (hμ : μ ≠ 0) (hf : ae_measurable f μ) : ∃ x, f x ≤ ⨍⁻ a, f a ∂μ := +let ⟨x, hx⟩ := nonempty_of_measure_ne_zero (measure_le_laverage_pos hμ hf).ne' in ⟨x, hx⟩ + +/-- **First moment method**. The minimum of a measurable function is smaller than its mean, while +avoiding a null set. -/ +lemma exists_not_mem_null_le_laverage (hμ : μ ≠ 0) (hf : ae_measurable f μ) (hN : μ N = 0) : + ∃ x ∉ N, f x ≤ ⨍⁻ a, f a ∂μ := +begin + have := measure_le_laverage_pos hμ hf, + rw ←measure_diff_null hN at this, + obtain ⟨x, hx, hxN⟩ := nonempty_of_measure_ne_zero this.ne', + exact ⟨x, hxN, hx⟩, +end + +end finite_measure + +section probability_measure +variables [is_probability_measure μ] + +/-- **First moment method**. A measurable function is smaller than its integral on a set of +positive measure. -/ +lemma measure_le_lintegral_pos (hf : ae_measurable f μ) : 0 < μ {x | f x ≤ ∫⁻ a, f a ∂μ} := +by simpa only [laverage_eq_lintegral] + using measure_le_laverage_pos (is_probability_measure.ne_zero μ) hf + +/-- **First moment method**. A measurable function is greater than its integral on a set of +positive measure. -/ +lemma measure_lintegral_le_pos (hint : ∫⁻ a, f a ∂μ ≠ ∞) : 0 < μ {x | ∫⁻ a, f a ∂μ ≤ f x} := +by simpa only [laverage_eq_lintegral] + using measure_laverage_le_pos (is_probability_measure.ne_zero μ) hint + +/-- **First moment method**. The minimum of a measurable function is smaller than its integral. -/ +lemma exists_le_lintegral (hf : ae_measurable f μ) : ∃ x, f x ≤ ∫⁻ a, f a ∂μ := +by simpa only [laverage_eq_lintegral] + using exists_le_laverage (is_probability_measure.ne_zero μ) hf + +/-- **First moment method**. The maximum of a measurable function is greater than its integral. -/ +lemma exists_lintegral_le (hint : ∫⁻ a, f a ∂μ ≠ ∞) : ∃ x, ∫⁻ a, f a ∂μ ≤ f x := +by simpa only [laverage_eq_lintegral] + using exists_laverage_le (is_probability_measure.ne_zero μ) hint + +/-- **First moment method**. The minimum of a measurable function is smaller than its integral, +while avoiding a null set. -/ +lemma exists_not_mem_null_le_lintegral (hf : ae_measurable f μ) (hN : μ N = 0) : + ∃ x ∉ N, f x ≤ ∫⁻ a, f a ∂μ := +by simpa only [laverage_eq_lintegral] + using exists_not_mem_null_le_laverage (is_probability_measure.ne_zero μ) hf hN + +/-- **First moment method**. The maximum of a measurable function is greater than its integral, +while avoiding a null set. -/ +lemma exists_not_mem_null_lintegral_le (hint : ∫⁻ a, f a ∂μ ≠ ∞) + (hN : μ N = 0) : ∃ x ∉ N, ∫⁻ a, f a ∂μ ≤ f x := +by simpa only [laverage_eq_lintegral] + using exists_not_mem_null_laverage_le (is_probability_measure.ne_zero μ) hint hN + +end probability_measure +end first_moment_ennreal end measure_theory diff --git a/src/measure_theory/integral/lebesgue.lean b/src/measure_theory/integral/lebesgue.lean index 0af668bac16f0..6041f5c3b1747 100644 --- a/src/measure_theory/integral/lebesgue.lean +++ b/src/measure_theory/integral/lebesgue.lean @@ -83,7 +83,7 @@ section lintegral open simple_func -variables {m : measurable_space α} {μ ν : measure α} +variables {m : measurable_space α} {μ ν : measure α} {f : α → ℝ≥0∞} {s : set α} /-- The **lower Lebesgue integral** of a function `f` with respect to a measure `μ`. -/ @[irreducible] def lintegral {m : measurable_space α} (μ : measure α) (f : α → ℝ≥0∞) : ℝ≥0∞ := @@ -758,6 +758,9 @@ begin exact hf (measurable_set_singleton r) end +@[simp] lemma lintegral_indicator_one (hs : measurable_set s) : ∫⁻ a, s.indicator 1 a ∂μ = μ s := +(lintegral_indicator_const hs _).trans $ one_mul _ + /-- A version of **Markov's inequality** for two functions. It doesn't follow from the standard Markov's inequality because we only assume measurability of `g`, not `f`. -/ lemma lintegral_add_mul_meas_add_le_le_lintegral {f g : α → ℝ≥0∞} (hle : f ≤ᵐ[μ] g) @@ -792,12 +795,25 @@ lemma mul_meas_ge_le_lintegral {f : α → ℝ≥0∞} (hf : measurable f) (ε : ε * μ {x | ε ≤ f x} ≤ ∫⁻ a, f a ∂μ := mul_meas_ge_le_lintegral₀ hf.ae_measurable ε -lemma lintegral_eq_top_of_measure_eq_top_pos {f : α → ℝ≥0∞} (hf : ae_measurable f μ) - (hμf : 0 < μ {x | f x = ∞}) : ∫⁻ x, f x ∂μ = ∞ := +lemma lintegral_eq_top_of_measure_eq_top_ne_zero {f : α → ℝ≥0∞} (hf : ae_measurable f μ) + (hμf : μ {x | f x = ∞} ≠ 0) : ∫⁻ x, f x ∂μ = ∞ := eq_top_iff.mpr $ -calc ∞ = ∞ * μ {x | ∞ ≤ f x} : by simp [mul_eq_top, hμf.ne.symm] +calc ∞ = ∞ * μ {x | ∞ ≤ f x} : by simp [mul_eq_top, hμf] ... ≤ ∫⁻ x, f x ∂μ : mul_meas_ge_le_lintegral₀ hf ∞ +lemma set_lintegral_eq_top_of_measure_eq_top_ne_zero (hf : ae_measurable f (μ.restrict s)) + (hμf : μ {x ∈ s | f x = ∞} ≠ 0) : ∫⁻ x in s, f x ∂μ = ∞ := +lintegral_eq_top_of_measure_eq_top_ne_zero hf $ + mt (eq_bot_mono $ by { rw ←set_of_inter_eq_sep, exact measure.le_restrict_apply _ _ }) hμf + +lemma measure_eq_top_of_lintegral_ne_top (hf : ae_measurable f μ) (hμf : ∫⁻ x, f x ∂μ ≠ ∞) : + μ {x | f x = ∞} = 0 := +of_not_not $ λ h, hμf $ lintegral_eq_top_of_measure_eq_top_ne_zero hf h + +lemma measure_eq_top_of_set_lintegral_ne_top (hf : ae_measurable f (μ.restrict s)) + (hμf : ∫⁻ x in s, f x ∂μ ≠ ∞) : μ {x ∈ s | f x = ∞} = 0 := +of_not_not $ λ h, hμf $ set_lintegral_eq_top_of_measure_eq_top_ne_zero hf h + /-- **Markov's inequality** also known as **Chebyshev's first inequality**. -/ lemma meas_ge_le_lintegral_div {f : α → ℝ≥0∞} (hf : ae_measurable f μ) {ε : ℝ≥0∞} (hε : ε ≠ 0) (hε' : ε ≠ ∞) : diff --git a/src/measure_theory/measurable_space.lean b/src/measure_theory/measurable_space.lean index 60e479ef9cf97..98a586b1589a6 100644 --- a/src/measure_theory/measurable_space.lean +++ b/src/measure_theory/measurable_space.lean @@ -191,7 +191,7 @@ lemma measurable_of_subsingleton_codomain [subsingleton β] (f : α → β) : measurable f := λ s hs, subsingleton.set_cases measurable_set.empty measurable_set.univ s -@[to_additive] +@[measurability, to_additive] lemma measurable_one [has_one α] : measurable (1 : β → α) := @measurable_const _ _ _ _ 1 lemma measurable_of_empty [is_empty α] (f : α → β) : measurable f := diff --git a/src/measure_theory/measure/measure_space_def.lean b/src/measure_theory/measure/measure_space_def.lean index cb8949247984b..1669f42d69e17 100644 --- a/src/measure_theory/measure/measure_space_def.lean +++ b/src/measure_theory/measure/measure_space_def.lean @@ -442,9 +442,9 @@ lemma inter_ae_eq_empty_of_ae_eq_empty_right (h : t =ᵐ[μ] (∅ : set α)) : by { convert ae_eq_set_inter (ae_eq_refl s) h, rw inter_empty, } @[to_additive] -lemma _root_.set.mul_indicator_ae_eq_one {M : Type*} [has_one M] {f : α → M} {s : set α} - (h : s.mul_indicator f =ᵐ[μ] 1) : μ (s ∩ function.mul_support f) = 0 := -by simpa [filter.eventually_eq, ae_iff] using h +lemma _root_.set.mul_indicator_ae_eq_one {M : Type*} [has_one M] {f : α → M} {s : set α} : + s.mul_indicator f =ᵐ[μ] 1 ↔ μ (s ∩ f.mul_support) = 0 := +by simpa [eventually_eq, eventually_iff, measure.ae, compl_set_of] /-- If `s ⊆ t` modulo a set of measure `0`, then `μ s ≤ μ t`. -/ @[mono] lemma measure_mono_ae (H : s ≤ᵐ[μ] t) : μ s ≤ μ t := diff --git a/src/probability/density.lean b/src/probability/density.lean index a4c7666e7d7e5..c2bf0bcb2b140 100644 --- a/src/probability/density.lean +++ b/src/probability/density.lean @@ -341,7 +341,7 @@ begin simp [hnt] }, rw [heq, set.inter_univ] at this, exact hns this }, - exact set.indicator_ae_eq_zero hu.symm, + exact set.indicator_ae_eq_zero.1 hu.symm, end lemma pdf_to_real_ae_eq {m : measurable_space Ω} From e5ab837fc252451f3eb9124ae6e7b6f57455e7b9 Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Wed, 21 Jun 2023 09:24:30 +0000 Subject: [PATCH 1208/1291] feat(manifold/{cont_mdiff,algebra/smooth_functions}): restriction as a homomorphism between spaces of smooth functions (#19209) This splits out for separate review the part of #19094 that is blocking the port. --- .../manifold/algebra/smooth_functions.lean | 60 +++++++++++++++++++ src/geometry/manifold/cont_mdiff.lean | 22 +++++++ 2 files changed, 82 insertions(+) diff --git a/src/geometry/manifold/algebra/smooth_functions.lean b/src/geometry/manifold/algebra/smooth_functions.lean index fe18fe42b2c32..932b3e776ebf1 100644 --- a/src/geometry/manifold/algebra/smooth_functions.lean +++ b/src/geometry/manifold/algebra/smooth_functions.lean @@ -15,6 +15,7 @@ In this file, we define instances of algebraic structures over smooth functions. noncomputable theory open_locale manifold +open topological_space variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] {E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] @@ -85,6 +86,38 @@ def coe_fn_monoid_hom {G : Type*} [monoid G] [topological_space G] [charted_space H' G] [has_smooth_mul I' G] : C^∞⟮I, N; I', G⟯ →* (N → G) := { to_fun := coe_fn, map_one' := coe_one, map_mul' := coe_mul } +variables (I N) + +/-- For a manifold `N` and a smooth homomorphism `φ` between Lie groups `G'`, `G''`, the +'left-composition-by-`φ`' group homomorphism from `C^∞⟮I, N; I', G'⟯` to `C^∞⟮I, N; I'', G''⟯`. -/ +@[to_additive "For a manifold `N` and a smooth homomorphism `φ` between additive Lie groups `G'`, +`G''`, the 'left-composition-by-`φ`' group homomorphism from `C^∞⟮I, N; I', G'⟯` to +`C^∞⟮I, N; I'', G''⟯`."] +def comp_left_monoid_hom + {G' : Type*} [monoid G'] [topological_space G'] [charted_space H' G'] [has_smooth_mul I' G'] + {G'' : Type*} [monoid G''] [topological_space G''] [charted_space H'' G''] + [has_smooth_mul I'' G''] (φ : G' →* G'') (hφ : smooth I' I'' φ) : + C^∞⟮I, N; I', G'⟯ →* C^∞⟮I, N; I'', G''⟯ := +{ to_fun := λ f, ⟨φ ∘ f, λ x, (hφ.smooth _).comp x (f.cont_mdiff x)⟩, + map_one' := by ext x; show φ 1 = 1; simp, + map_mul' := λ f g, by ext x; show φ (f x * g x) = φ (f x) * φ (g x); simp } + +variables (I') {N} + +/-- For a Lie group `G` and open sets `U ⊆ V` in `N`, the 'restriction' group homomorphism from +`C^∞⟮I, V; I', G⟯` to `C^∞⟮I, U; I', G⟯`. -/ +@[to_additive "For an additive Lie group `G` and open sets `U ⊆ V` in `N`, the 'restriction' group +homomorphism from `C^∞⟮I, V; I', G⟯` to `C^∞⟮I, U; I', G⟯`."] +def restrict_monoid_hom + (G : Type*) [monoid G] [topological_space G] [charted_space H' G] [has_smooth_mul I' G] + {U V : opens N} (h : U ≤ V) : + C^∞⟮I, V; I', G⟯ →* C^∞⟮I, U; I', G⟯ := +{ to_fun := λ f, ⟨f ∘ set.inclusion h, f.smooth.comp (smooth_inclusion h)⟩, + map_one' := rfl, + map_mul' := λ f g, rfl } + +variables {I N I' N'} + @[to_additive] instance comm_monoid {G : Type*} [comm_monoid G] [topological_space G] [charted_space H' G] [has_smooth_mul I' G] : @@ -155,6 +188,33 @@ instance comm_ring {R : Type*} [comm_ring R] [topological_space R] ..smooth_map.add_comm_group, ..smooth_map.comm_monoid,} +variables (I N) + +/-- For a manifold `N` and a smooth homomorphism `φ` between smooth rings `R'`, `R''`, the +'left-composition-by-`φ`' ring homomorphism from `C^∞⟮I, N; I', R'⟯` to `C^∞⟮I, N; I'', R''⟯`. -/ +def comp_left_ring_hom + {R' : Type*} [ring R'] [topological_space R'] [charted_space H' R'] [smooth_ring I' R'] + {R'' : Type*} [ring R''] [topological_space R''] [charted_space H'' R''] [smooth_ring I'' R''] + (φ : R' →+* R'') (hφ : smooth I' I'' φ) : + C^∞⟮I, N; I', R'⟯ →+* C^∞⟮I, N; I'', R''⟯ := +{ to_fun := λ f, ⟨φ ∘ f, λ x, (hφ.smooth _).comp x (f.cont_mdiff x)⟩, + .. smooth_map.comp_left_monoid_hom I N φ.to_monoid_hom hφ, + .. smooth_map.comp_left_add_monoid_hom I N φ.to_add_monoid_hom hφ } + +variables (I') {N} + +/-- For a "smooth ring" `R` and open sets `U ⊆ V` in `N`, the "restriction" ring homomorphism from +`C^∞⟮I, V; I', R⟯` to `C^∞⟮I, U; I', R⟯`. -/ +def restrict_ring_hom + (R : Type*) [ring R] [topological_space R] [charted_space H' R] [smooth_ring I' R] + {U V : opens N} (h : U ≤ V) : + C^∞⟮I, V; I', R⟯ →+* C^∞⟮I, U; I', R⟯ := +{ to_fun := λ f, ⟨f ∘ set.inclusion h, f.smooth.comp (smooth_inclusion h)⟩, + .. smooth_map.restrict_monoid_hom I I' R h, + .. smooth_map.restrict_add_monoid_hom I I' R h } + +variables {I N I' N'} + /-- Coercion to a function as a `ring_hom`. -/ @[simps] def coe_fn_ring_hom {R : Type*} [comm_ring R] [topological_space R] diff --git a/src/geometry/manifold/cont_mdiff.lean b/src/geometry/manifold/cont_mdiff.lean index 72c43255f6672..e6205468df232 100644 --- a/src/geometry/manifold/cont_mdiff.lean +++ b/src/geometry/manifold/cont_mdiff.lean @@ -1245,6 +1245,28 @@ begin exact cont_mdiff_at_const } end +/-! ### The inclusion map from one open set to another is smooth -/ +section +open topological_space + +lemma cont_mdiff_inclusion {n : ℕ∞} {U V : opens M} (h : U ≤ V) : + cont_mdiff I I n (set.inclusion h : U → V) := +begin + rintros ⟨x, hx : x ∈ U⟩, + apply (cont_diff_within_at_local_invariant_prop I I n).lift_prop_inclusion, + intros y, + dsimp [cont_diff_within_at_prop], + rw [set.univ_inter], + refine cont_diff_within_at_id.congr _ _, + { exact I.right_inv_on }, + { exact congr_arg I (I.left_inv y) }, +end + +lemma smooth_inclusion {U V : opens M} (h : U ≤ V) : smooth I I (set.inclusion h : U → V) := +cont_mdiff_inclusion h + +end + /-! ### Equivalence with the basic definition for functions between vector spaces -/ section module From f23a09ce6d3f367220dc3cecad6b7eb69eb01690 Mon Sep 17 00:00:00 2001 From: Aaron Bies <48803074+slerpyyy@users.noreply.github.com> Date: Wed, 21 Jun 2023 15:57:07 +0000 Subject: [PATCH 1209/1291] feat(analysis/special_functions/log/base): add `real.logb_mul_base` (#18979) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Add proof that states `logb (a * b) c = ((logb a c)⁻¹ + (logb b c)⁻¹)⁻¹`. Co-authored-by: Eric Wieser --- src/analysis/special_functions/log/base.lean | 33 +++++++++++++++++++ src/analysis/special_functions/log/basic.lean | 3 ++ 2 files changed, 36 insertions(+) diff --git a/src/analysis/special_functions/log/base.lean b/src/analysis/special_functions/log/base.lean index 02b23f546daa7..b7a3e10678ea8 100644 --- a/src/analysis/special_functions/log/base.lean +++ b/src/analysis/special_functions/log/base.lean @@ -55,6 +55,39 @@ by simp_rw [logb, log_div hx hy, sub_div] @[simp] lemma logb_inv (x : ℝ) : logb b (x⁻¹) = -logb b x := by simp [logb, neg_div] +lemma inv_logb (a b : ℝ) : (logb a b)⁻¹ = logb b a := by simp_rw [logb, inv_div] + +theorem inv_logb_mul_base {a b : ℝ} (h₁ : a ≠ 0) (h₂ : b ≠ 0) (c : ℝ) : + (logb (a * b) c)⁻¹ = (logb a c)⁻¹ + (logb b c)⁻¹ := +by simp_rw inv_logb; exact logb_mul h₁ h₂ + +theorem inv_logb_div_base {a b : ℝ} (h₁ : a ≠ 0) (h₂ : b ≠ 0) (c : ℝ) : + (logb (a / b) c)⁻¹ = (logb a c)⁻¹ - (logb b c)⁻¹ := +by simp_rw inv_logb; exact logb_div h₁ h₂ + +theorem logb_mul_base {a b : ℝ} (h₁ : a ≠ 0) (h₂ : b ≠ 0) (c : ℝ) : + logb (a * b) c = ((logb a c)⁻¹ + (logb b c)⁻¹)⁻¹ := +by rw [←inv_logb_mul_base h₁ h₂ c, inv_inv] + +theorem logb_div_base {a b : ℝ} (h₁ : a ≠ 0) (h₂ : b ≠ 0) (c : ℝ) : + logb (a / b) c = ((logb a c)⁻¹ - (logb b c)⁻¹)⁻¹ := +by rw [←inv_logb_div_base h₁ h₂ c, inv_inv] + +theorem mul_logb {a b c : ℝ} (h₁ : b ≠ 0) (h₂ : b ≠ 1) (h₃ : b ≠ -1) : + logb a b * logb b c = logb a c := +begin + unfold logb, + rw [mul_comm, div_mul_div_cancel _ (log_ne_zero.mpr ⟨h₁, h₂, h₃⟩)], +end + +theorem div_logb {a b c : ℝ} (h₁ : c ≠ 0) (h₂ : c ≠ 1) (h₃ : c ≠ -1) : + logb a c / logb b c = logb a b := +begin + unfold logb, + -- TODO: div_div_div_cancel_left is missing for `group_with_zero`, + rw [div_div_div_eq, mul_comm, mul_div_mul_right _ _ (log_ne_zero.mpr ⟨h₁, h₂, h₃⟩)], +end + section b_pos_and_ne_one variable (b_pos : 0 < b) diff --git a/src/analysis/special_functions/log/basic.lean b/src/analysis/special_functions/log/basic.lean index 1f0069f4f4472..da7392ce78e95 100644 --- a/src/analysis/special_functions/log/basic.lean +++ b/src/analysis/special_functions/log/basic.lean @@ -196,6 +196,9 @@ begin { rintro (rfl|rfl|rfl); simp only [log_one, log_zero, log_neg_eq_log], } end +lemma log_ne_zero {x : ℝ} : log x ≠ 0 ↔ x ≠ 0 ∧ x ≠ 1 ∧ x ≠ -1 := +by simpa only [not_or_distrib] using log_eq_zero.not + @[simp] lemma log_pow (x : ℝ) (n : ℕ) : log (x ^ n) = n * log x := begin induction n with n ih, From eba7871095e834365616b5e43c8c7bb0b37058d0 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 21 Jun 2023 18:29:27 +0000 Subject: [PATCH 1210/1291] =?UTF-8?q?feat(data/finset/pointwise):=20`|s|?= =?UTF-8?q?=20=E2=88=A3=20|s=20*=20t|`=20(#18663)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- src/data/finset/basic.lean | 8 ++++++++ src/data/finset/n_ary.lean | 31 +++++++++++++++++++++++++++++++ src/data/finset/pointwise.lean | 18 ++++++++++++++++++ 3 files changed, 57 insertions(+) diff --git a/src/data/finset/basic.lean b/src/data/finset/basic.lean index 30687beee91e0..43bca0b7aca26 100644 --- a/src/data/finset/basic.lean +++ b/src/data/finset/basic.lean @@ -570,6 +570,10 @@ lemma forall_mem_cons (h : a ∉ s) (p : α → Prop) : (∀ x, x ∈ cons a s h → p x) ↔ p a ∧ ∀ x, x ∈ s → p x := by simp only [mem_cons, or_imp_distrib, forall_and_distrib, forall_eq] +/-- Useful in proofs by induction. -/ +lemma forall_of_forall_cons {p : α → Prop} {h : a ∉ s} (H : ∀ x, x ∈ cons a s h → p x) (x) + (h : x ∈ s) : p x := H _ $ mem_cons.2 $ or.inr h + @[simp] lemma mk_cons {s : multiset α} (h : (a ::ₘ s).nodup) : (⟨a ::ₘ s, h⟩ : finset α) = cons a ⟨s, (nodup_cons.1 h).2⟩ (nodup_cons.1 h).1 := rfl @@ -2084,6 +2088,10 @@ lemma forall_mem_insert [decidable_eq α] (a : α) (s : finset α) (p : α → P (∀ x, x ∈ insert a s → p x) ↔ p a ∧ ∀ x, x ∈ s → p x := by simp only [mem_insert, or_imp_distrib, forall_and_distrib, forall_eq] +/-- Useful in proofs by induction. -/ +lemma forall_of_forall_insert [decidable_eq α] {p : α → Prop} {a : α} {s : finset α} + (H : ∀ x, x ∈ insert a s → p x) (x) (h : x ∈ s) : p x := H _ $ mem_insert_of_mem h + end finset /-- Equivalence between the set of natural numbers which are `≥ k` and `ℕ`, given by `n → n - k`. -/ diff --git a/src/data/finset/n_ary.lean b/src/data/finset/n_ary.lean index 3a0db131eab7b..dd7d4df71f086 100644 --- a/src/data/finset/n_ary.lean +++ b/src/data/finset/n_ary.lean @@ -358,6 +358,37 @@ lemma image₂_right_identity {f : γ → β → γ} {b : β} (h : ∀ a, f a b image₂ f s {b} = s := by rw [image₂_singleton_right, funext h, image_id'] +/-- If each partial application of `f` is injective, and images of `s` under those partial +applications are disjoint (but not necessarily distinct!), then the size of `t` divides the size of +`finset.image₂ f s t`. -/ +lemma card_dvd_card_image₂_right (hf : ∀ a ∈ s, injective (f a)) + (hs : ((λ a, t.image $ f a) '' s).pairwise_disjoint id) : + t.card ∣ (image₂ f s t).card := +begin + classical, + induction s using finset.induction with a s ha ih, + { simp }, + specialize ih (forall_of_forall_insert hf) + (hs.subset $ set.image_subset _ $ coe_subset.2 $ subset_insert _ _), + rw image₂_insert_left, + by_cases h : disjoint (image (f a) t) (image₂ f s t), + { rw card_union_eq h, + exact (card_image_of_injective _ $ hf _ $ mem_insert_self _ _).symm.dvd.add ih }, + simp_rw [←bUnion_image_left, disjoint_bUnion_right, not_forall] at h, + obtain ⟨b, hb, h⟩ := h, + rwa union_eq_right_iff_subset.2, + exact (hs.eq (set.mem_image_of_mem _ $ mem_insert_self _ _) + (set.mem_image_of_mem _ $ mem_insert_of_mem hb) h).trans_subset (image_subset_image₂_right hb), +end + +/-- If each partial application of `f` is injective, and images of `t` under those partial +applications are disjoint (but not necessarily distinct!), then the size of `s` divides the size of +`finset.image₂ f s t`. -/ +lemma card_dvd_card_image₂_left (hf : ∀ b ∈ t, injective (λ a, f a b)) + (ht : ((λ b, s.image $ λ a, f a b) '' t).pairwise_disjoint id) : + s.card ∣ (image₂ f s t).card := +by { rw ←image₂_swap, exact card_dvd_card_image₂_right hf ht } + variables [decidable_eq α] [decidable_eq β] lemma image₂_inter_union_subset_union : diff --git a/src/data/finset/pointwise.lean b/src/data/finset/pointwise.lean index 03063a1ed2bc6..0f441be5a12ec 100644 --- a/src/data/finset/pointwise.lean +++ b/src/data/finset/pointwise.lean @@ -1078,6 +1078,24 @@ coe_injective $ by { push_cast, exact set.smul_univ hs } @[simp, to_additive] lemma card_smul_finset (a : α) (s : finset β) : (a • s).card = s.card := card_image_of_injective _ $ mul_action.injective _ +/-- If the left cosets of `t` by elements of `s` are disjoint (but not necessarily distinct!), then +the size of `t` divides the size of `s * t`. -/ +@[to_additive "If the left cosets of `t` by elements of `s` are disjoint (but not necessarily +distinct!), then the size of `t` divides the size of `s + t`."] +lemma card_dvd_card_smul_right {s : finset α} : + ((• t) '' (s : set α)).pairwise_disjoint id → t.card ∣ (s • t).card := +card_dvd_card_image₂_right (λ _ _, mul_action.injective _) + +variables [decidable_eq α] + +/-- If the right cosets of `s` by elements of `t` are disjoint (but not necessarily distinct!), then +the size of `s` divides the size of `s * t`. -/ +@[to_additive "If the right cosets of `s` by elements of `t` are disjoint (but not necessarily +distinct!), then the size of `s` divides the size of `s + t`."] +lemma card_dvd_card_mul_left {s t : finset α} : + ((λ b, s.image $ λ a, a * b) '' (t : set α)).pairwise_disjoint id → s.card ∣ (s * t).card := +card_dvd_card_image₂_left (λ _ _, mul_left_injective _) + end group section group_with_zero From 3efd324a3a31eaa40c9d5bfc669c4fafee5f9423 Mon Sep 17 00:00:00 2001 From: Anatole Dedecker Date: Wed, 21 Jun 2023 18:29:28 +0000 Subject: [PATCH 1211/1291] feat(topology/algebra/order/compact): remove conditional completeness assumption in `is_compact.exists_forall_le` (#18991) --- src/order/directed.lean | 6 +- src/topology/algebra/order/compact.lean | 219 ++++++++++++++++-------- src/topology/order/basic.lean | 42 ----- src/topology/subset_properties.lean | 42 +++-- 4 files changed, 172 insertions(+), 137 deletions(-) diff --git a/src/order/directed.lean b/src/order/directed.lean index b5c5f44c6f8fc..83c44b4bf2b62 100644 --- a/src/order/directed.lean +++ b/src/order/directed.lean @@ -131,6 +131,10 @@ lemma directed_on_of_inf_mem [semilattice_inf α] {S : set α} (H : ∀ ⦃i j⦄, i ∈ S → j ∈ S → i ⊓ j ∈ S) : directed_on (≥) S := λ a ha b hb, ⟨a ⊓ b, H ha hb, inf_le_left, inf_le_right⟩ +lemma is_total.directed [is_total α r] (f : ι → α) : + directed r f := +λ i j, or.cases_on (total_of r (f i) (f j)) (λ h, ⟨j, h, refl _⟩) (λ h, ⟨i, refl _, h⟩) + /-- `is_directed α r` states that for any elements `a`, `b` there exists an element `c` such that `r a c` and `r b c`. -/ class is_directed (α : Type*) (r : α → α → Prop) : Prop := @@ -150,7 +154,7 @@ lemma directed_on_univ_iff : directed_on r set.univ ↔ is_directed α r := @[priority 100] -- see Note [lower instance priority] instance is_total.to_is_directed [is_total α r] : is_directed α r := -⟨λ a b, or.cases_on (total_of r a b) (λ h, ⟨b, h, refl _⟩) (λ h, ⟨a, refl _, h⟩)⟩ +by rw ← directed_id_iff; exact is_total.directed _ lemma is_directed_mono [is_directed α r] (h : ∀ ⦃a b⦄, r a b → s a b) : is_directed α s := ⟨λ a b, let ⟨c, ha, hb⟩ := is_directed.directed a b in ⟨c, h ha, h hb⟩⟩ diff --git a/src/topology/algebra/order/compact.lean b/src/topology/algebra/order/compact.lean index c7e46e18e122e..0eedd6364138d 100644 --- a/src/topology/algebra/order/compact.lean +++ b/src/topology/algebra/order/compact.lean @@ -136,80 +136,39 @@ is_compact_iff_compact_space.mp is_compact_Icc end /-! -### Min and max elements of a compact set +### Extreme value theorem -/ -variables {α β γ : Type*} [conditionally_complete_linear_order α] [topological_space α] - [order_topology α] [topological_space β] [topological_space γ] - -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 αᵒᵈ _ _ _ _ hs ne_s - -lemma is_compact.is_glb_Inf {s : set α} (hs : is_compact s) (ne_s : s.nonempty) : - is_glb s (Inf s) := -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 αᵒᵈ _ _ _ _ hs ne_s - -lemma is_compact.is_least_Inf {s : set α} (hs : is_compact s) (ne_s : s.nonempty) : - is_least s (Inf s) := -⟨hs.Inf_mem ne_s, (hs.is_glb_Inf ne_s).1⟩ +section linear_order -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 αᵒᵈ _ _ _ _ hs ne_s +variables {α β γ : Type*} [linear_order α] [topological_space α] + [order_closed_topology α] [topological_space β] [topological_space γ] lemma is_compact.exists_is_least {s : set α} (hs : is_compact s) (ne_s : s.nonempty) : ∃ x, is_least s x := -⟨_, hs.is_least_Inf ne_s⟩ +begin + haveI : nonempty s := ne_s.to_subtype, + suffices : (s ∩ ⋂ x ∈ s, Iic x).nonempty, + from ⟨this.some, this.some_spec.1, mem_Inter₂.mp this.some_spec.2⟩, + rw bInter_eq_Inter, + by_contra H, + rw not_nonempty_iff_eq_empty at H, + rcases hs.elim_directed_family_closed (λ x : s, Iic ↑x) (λ x, is_closed_Iic) H + ((is_total.directed coe).mono_comp _ (λ _ _, Iic_subset_Iic.mpr)) with ⟨x, hx⟩, + exact not_nonempty_iff_eq_empty.mpr hx ⟨x, x.2, le_rfl⟩ +end lemma is_compact.exists_is_greatest {s : set α} (hs : is_compact s) (ne_s : s.nonempty) : ∃ x, is_greatest s x := -⟨_, hs.is_greatest_Sup ne_s⟩ +@is_compact.exists_is_least αᵒᵈ _ _ _ _ hs ne_s lemma is_compact.exists_is_glb {s : set α} (hs : is_compact s) (ne_s : s.nonempty) : ∃ x ∈ s, is_glb s x := -⟨_, hs.Inf_mem ne_s, hs.is_glb_Inf ne_s⟩ +exists_imp_exists (λ x (hx : is_least s x), ⟨hx.1, hx.is_glb⟩) (hs.exists_is_least ne_s) lemma is_compact.exists_is_lub {s : set α} (hs : is_compact s) (ne_s : s.nonempty) : ∃ x ∈ s, is_lub s x := -⟨_, hs.Sup_mem ne_s, hs.is_lub_Sup ne_s⟩ - -lemma is_compact.exists_Inf_image_eq_and_le {s : set β} (hs : is_compact s) (ne_s : s.nonempty) - {f : β → α} (hf : continuous_on f s) : - ∃ x ∈ s, Inf (f '' s) = f x ∧ ∀ y ∈ s, f x ≤ f y := -let ⟨x, hxs, hx⟩ := (hs.image_of_continuous_on hf).Inf_mem (ne_s.image f) -in ⟨x, hxs, hx.symm, λ y hy, - hx.trans_le $ cInf_le (hs.image_of_continuous_on hf).bdd_below $ mem_image_of_mem f 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 αᵒᵈ _ _ _ _ _ _ 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) : - ∃ x ∈ s, Inf (f '' s) = f x := -let ⟨x, hxs, hx, _⟩ := hs.exists_Inf_image_eq_and_le ne_s hf in ⟨x, hxs, hx⟩ - -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 αᵒᵈ _ _ _ _ _ - -lemma eq_Icc_of_connected_compact {s : set α} (h₁ : is_connected s) (h₂ : is_compact s) : - s = Icc (Inf s) (Sup s) := -eq_Icc_cInf_cSup_of_connected_bdd_closed h₁ h₂.bdd_below h₂.bdd_above h₂.is_closed - -/-! -### Extreme value theorem --/ +@is_compact.exists_is_glb αᵒᵈ _ _ _ _ hs ne_s /-- The **extreme value theorem**: a continuous function realizes its minimum on a compact set. -/ lemma is_compact.exists_forall_le {s : set β} (hs : is_compact s) (ne_s : s.nonempty) @@ -277,6 +236,68 @@ lemma continuous.exists_forall_ge [nonempty β] {f : β → α} ∃ x, ∀ y, f y ≤ f x := @continuous.exists_forall_le αᵒᵈ _ _ _ _ _ _ _ hf hlim +/-- A continuous function with compact support has a global minimum. -/ +@[to_additive "A continuous function with compact support has a global minimum."] +lemma continuous.exists_forall_le_of_has_compact_mul_support [nonempty β] [has_one α] + {f : β → α} (hf : continuous f) (h : has_compact_mul_support f) : + ∃ (x : β), ∀ (y : β), f x ≤ f y := +begin + 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. -/ +@[to_additive "A continuous function with compact support has a global maximum."] +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 αᵒᵈ _ _ _ _ _ _ _ _ hf h + +/-- A compact set is bounded below -/ +lemma is_compact.bdd_below [nonempty α] {s : set α} (hs : is_compact s) : bdd_below s := +begin + cases s.eq_empty_or_nonempty, + { rw h, + exact bdd_below_empty }, + { obtain ⟨a, ha, has⟩ := hs.exists_is_least h, + exact ⟨a, has⟩ }, +end + +/-- A compact set is bounded above -/ +lemma is_compact.bdd_above [nonempty α] {s : set α} (hs : is_compact s) : bdd_above s := +@is_compact.bdd_below αᵒᵈ _ _ _ _ _ hs + +/-- A continuous function is bounded below on a compact set. -/ +lemma is_compact.bdd_below_image [nonempty α] {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 [nonempty α] {f : β → α} {K : set β} + (hK : is_compact K) (hf : continuous_on f K) : bdd_above (f '' K) := +@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. "-/] +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 αᵒᵈ _ _ _ _ _ _ _ hf h + +end linear_order + +section conditionally_complete_linear_order + +variables {α β γ : Type*} [conditionally_complete_linear_order α] [topological_space α] + [order_closed_topology α] [topological_space β] [topological_space γ] + lemma is_compact.Sup_lt_iff_of_continuous {f : β → α} {K : set β} (hK : is_compact K) (h0K : K.nonempty) (hf : continuous_on f K) (y : α) : Sup (f '' K) < y ↔ ∀ x ∈ K, f x < y := @@ -294,23 +315,71 @@ lemma is_compact.lt_Inf_iff_of_continuous {α β : Type*} y < Inf (f '' K) ↔ ∀ x ∈ K, y < f x := @is_compact.Sup_lt_iff_of_continuous αᵒᵈ β _ _ _ _ _ _ hK h0K hf y -/-- A continuous function with compact support has a global minimum. -/ -@[to_additive "A continuous function with compact support has a global minimum."] -lemma continuous.exists_forall_le_of_has_compact_mul_support [nonempty β] [has_one α] - {f : β → α} (hf : continuous f) (h : has_compact_mul_support f) : - ∃ (x : β), ∀ (y : β), f x ≤ f y := -begin - 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 +end conditionally_complete_linear_order -/-- A continuous function with compact support has a global maximum. -/ -@[to_additive "A continuous function with compact support has a global maximum."] -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 αᵒᵈ _ _ _ _ _ _ _ _ hf h +/-! +### Min and max elements of a compact set +-/ + +section order_closed_topology + +variables {α β γ : Type*} [conditionally_complete_linear_order α] [topological_space α] + [order_closed_topology α] [topological_space β] [topological_space γ] + +lemma is_compact.Inf_mem {s : set α} (hs : is_compact s) (ne_s : s.nonempty) : + Inf s ∈ s := +let ⟨a, ha⟩ := hs.exists_is_least ne_s in +ha.Inf_mem + +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) := +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 αᵒᵈ _ _ _ _ hs ne_s + +lemma is_compact.is_least_Inf {s : set α} (hs : is_compact s) (ne_s : s.nonempty) : + is_least s (Inf s) := +⟨hs.Inf_mem ne_s, (hs.is_glb_Inf ne_s).1⟩ + +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 αᵒᵈ _ _ _ _ hs ne_s + +lemma is_compact.exists_Inf_image_eq_and_le {s : set β} (hs : is_compact s) (ne_s : s.nonempty) + {f : β → α} (hf : continuous_on f s) : + ∃ x ∈ s, Inf (f '' s) = f x ∧ ∀ y ∈ s, f x ≤ f y := +let ⟨x, hxs, hx⟩ := (hs.image_of_continuous_on hf).Inf_mem (ne_s.image f) +in ⟨x, hxs, hx.symm, λ y hy, + hx.trans_le $ cInf_le (hs.image_of_continuous_on hf).bdd_below $ mem_image_of_mem f 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 αᵒᵈ _ _ _ _ _ _ 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) : + ∃ x ∈ s, Inf (f '' s) = f x := +let ⟨x, hxs, hx, _⟩ := hs.exists_Inf_image_eq_and_le ne_s hf in ⟨x, hxs, hx⟩ + +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 αᵒᵈ _ _ _ _ _ + +end order_closed_topology + +variables {α β γ : Type*} [conditionally_complete_linear_order α] [topological_space α] + [order_topology α] [topological_space β] [topological_space γ] + +lemma eq_Icc_of_connected_compact {s : set α} (h₁ : is_connected s) (h₂ : is_compact s) : + s = Icc (Inf s) (Sup s) := +eq_Icc_cInf_cSup_of_connected_bdd_closed h₁ h₂.bdd_below h₂.bdd_above h₂.is_closed lemma is_compact.continuous_Sup {f : γ → β → α} {K : set β} (hK : is_compact K) (hf : continuous ↿f) : diff --git a/src/topology/order/basic.lean b/src/topology/order/basic.lean index 733fa1a9de319..a17106de53df5 100644 --- a/src/topology/order/basic.lean +++ b/src/topology/order/basic.lean @@ -648,48 +648,6 @@ 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 αᵒᵈ _ _ _ _ _ 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 αᵒᵈ _ _ _ _ _ _ _ _ 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 αᵒᵈ _ _ _ _ _ _ _ _ hf h - end linear_order end order_closed_topology diff --git a/src/topology/subset_properties.lean b/src/topology/subset_properties.lean index 36f63e1120bbb..97d583e2dacd9 100644 --- a/src/topology/subset_properties.lean +++ b/src/topology/subset_properties.lean @@ -230,6 +230,19 @@ lemma is_compact.disjoint_nhds_set_right {l : filter α} (hs : is_compact s) : disjoint l (𝓝ˢ s) ↔ ∀ x ∈ s, disjoint l (𝓝 x) := by simpa only [disjoint.comm] using hs.disjoint_nhds_set_left +/-- For every directed family of closed sets whose intersection avoids a compact set, +there exists a single element of the family which itself avoids this compact set. -/ +lemma is_compact.elim_directed_family_closed {ι : Type v} [hι : nonempty ι] (hs : is_compact s) + (Z : ι → set α) (hZc : ∀ i, is_closed (Z i)) (hsZ : s ∩ (⋂ i, Z i) = ∅) (hdZ : directed (⊇) Z) : + ∃ i : ι, s ∩ Z i = ∅ := +let ⟨t, ht⟩ := hs.elim_directed_cover (compl ∘ Z) (λ i, (hZc i).is_open_compl) + (by simpa only [subset_def, not_forall, eq_empty_iff_forall_not_mem, mem_Union, + exists_prop, mem_inter_iff, not_and, iff_self, mem_Inter, mem_compl_iff] using hsZ) + (hdZ.mono_comp _ $ λ _ _, compl_subset_compl.mpr) + in +⟨t, by simpa only [subset_def, not_forall, eq_empty_iff_forall_not_mem, mem_Union, + exists_prop, mem_inter_iff, not_and, iff_self, mem_Inter, mem_compl_iff] using ht⟩ + /-- For every family of closed sets whose intersection avoids a compact set, there exists a finite subfamily whose intersection avoids this compact set. -/ lemma is_compact.elim_finite_subfamily_closed {s : set α} {ι : Type v} (hs : is_compact s) @@ -273,25 +286,16 @@ lemma is_compact.nonempty_Inter_of_directed_nonempty_compact_closed (hZn : ∀ i, (Z i).nonempty) (hZc : ∀ i, is_compact (Z i)) (hZcl : ∀ i, is_closed (Z i)) : (⋂ i, Z i).nonempty := begin - apply hι.elim, - intro i₀, - let Z' := λ i, Z i ∩ Z i₀, - suffices : (⋂ i, Z' i).nonempty, - { exact this.mono (Inter_mono $ λ i, inter_subset_left (Z i) (Z i₀)) }, - rw nonempty_iff_ne_empty, - intro H, - obtain ⟨t, ht⟩ : ∃ (t : finset ι), ((Z i₀) ∩ ⋂ (i ∈ t), Z' i) = ∅, - from (hZc i₀).elim_finite_subfamily_closed Z' - (assume i, is_closed.inter (hZcl i) (hZcl i₀)) (by rw [H, inter_empty]), - obtain ⟨i₁, hi₁⟩ : ∃ i₁ : ι, Z i₁ ⊆ Z i₀ ∧ ∀ i ∈ t, Z i₁ ⊆ Z' i, - { rcases directed.finset_le hZd t with ⟨i, hi⟩, - rcases hZd i i₀ with ⟨i₁, hi₁, hi₁₀⟩, - use [i₁, hi₁₀], - intros j hj, - exact subset_inter (subset.trans hi₁ (hi j hj)) hi₁₀ }, - suffices : ((Z i₀) ∩ ⋂ (i ∈ t), Z' i).nonempty, - { rw nonempty_iff_ne_empty at this, contradiction }, - exact (hZn i₁).mono (subset_inter hi₁.left $ subset_Inter₂ hi₁.right), + let i₀ := hι.some, + suffices : (Z i₀ ∩ ⋂ i, Z i).nonempty, + by rwa inter_eq_right_iff_subset.mpr (Inter_subset _ i₀) at this, + simp only [nonempty_iff_ne_empty] at hZn ⊢, + apply mt ((hZc i₀).elim_directed_family_closed Z hZcl), + push_neg, + simp only [← nonempty_iff_ne_empty] at hZn ⊢, + refine ⟨hZd, λ i, _⟩, + rcases hZd i₀ i with ⟨j, hji₀, hji⟩, + exact (hZn j).mono (subset_inter hji₀ hji) end /-- Cantor's intersection theorem for sequences indexed by `ℕ`: From 6285167a053ad0990fc88e56c48ccd9fae6550eb Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 21 Jun 2023 18:29:29 +0000 Subject: [PATCH 1212/1291] chore(topology/algebra/module/basic): generalize `to_span_singleton` to topological spaces (#19116) This should be straightforward to forward-port, as it involves deleting one hunk and pasting in another. --- .../normed_space/continuous_linear_map.lean | 21 ------------- src/topology/algebra/module/basic.lean | 31 +++++++++++++++++++ 2 files changed, 31 insertions(+), 21 deletions(-) diff --git a/src/analysis/normed_space/continuous_linear_map.lean b/src/analysis/normed_space/continuous_linear_map.lean index 693adffcd123e..87ad5e2c41002 100644 --- a/src/analysis/normed_space/continuous_linear_map.lean +++ b/src/analysis/normed_space/continuous_linear_map.lean @@ -192,27 +192,6 @@ lemma to_span_singleton_homothety (x : E) (c : 𝕜) : ‖linear_map.to_span_singleton 𝕜 E x c‖ = ‖x‖ * ‖c‖ := by {rw mul_comm, exact norm_smul _ _} -/-- Given an element `x` of a normed space `E` over a field `𝕜`, the natural continuous - linear map from `𝕜` to `E` by taking multiples of `x`.-/ -def to_span_singleton (x : E) : 𝕜 →L[𝕜] E := -of_homothety (linear_map.to_span_singleton 𝕜 E x) ‖x‖ (to_span_singleton_homothety 𝕜 x) - -lemma to_span_singleton_apply (x : E) (r : 𝕜) : to_span_singleton 𝕜 x r = r • x := -by simp [to_span_singleton, of_homothety, linear_map.to_span_singleton] - -lemma to_span_singleton_add (x y : E) : - to_span_singleton 𝕜 (x + y) = to_span_singleton 𝕜 x + to_span_singleton 𝕜 y := -by { ext1, simp [to_span_singleton_apply], } - -lemma to_span_singleton_smul' (𝕜') [normed_field 𝕜'] [normed_space 𝕜' E] - [smul_comm_class 𝕜 𝕜' E] (c : 𝕜') (x : E) : - to_span_singleton 𝕜 (c • x) = c • to_span_singleton 𝕜 x := -by { ext1, rw [to_span_singleton_apply, smul_apply, to_span_singleton_apply, smul_comm], } - -lemma to_span_singleton_smul (c : 𝕜) (x : E) : - to_span_singleton 𝕜 (c • x) = c • to_span_singleton 𝕜 x := -to_span_singleton_smul' 𝕜 𝕜 c x - end continuous_linear_map section diff --git a/src/topology/algebra/module/basic.lean b/src/topology/algebra/module/basic.lean index 5379fd79c140b..f7d3e1befccef 100644 --- a/src/topology/algebra/module/basic.lean +++ b/src/topology/algebra/module/basic.lean @@ -1003,6 +1003,37 @@ lemma smul_right_comp [has_continuous_mul R₁] {x : M₂} {c : R₁} : smul_right (1 : R₁ →L[R₁] R₁) (c • x) := by { ext, simp [mul_smul] } +section to_span_singleton +variables (R₁) +variables [has_continuous_smul R₁ M₁] + +/-- Given an element `x` of a topological space `M` over a semiring `R`, the natural continuous +linear map from `R` to `M` by taking multiples of `x`.-/ +def to_span_singleton (x : M₁) : R₁ →L[R₁] M₁ := +{ to_linear_map := linear_map.to_span_singleton R₁ M₁ x, + cont := continuous_id.smul continuous_const } + +lemma to_span_singleton_apply (x : M₁) (r : R₁) : to_span_singleton R₁ x r = r • x := +rfl + +lemma to_span_singleton_add [has_continuous_add M₁] (x y : M₁) : + to_span_singleton R₁ (x + y) = to_span_singleton R₁ x + to_span_singleton R₁ y := +by { ext1, simp [to_span_singleton_apply], } + +lemma to_span_singleton_smul' {α} [monoid α] [distrib_mul_action α M₁] + [has_continuous_const_smul α M₁] + [smul_comm_class R₁ α M₁] (c : α) (x : M₁) : + to_span_singleton R₁ (c • x) = c • to_span_singleton R₁ x := +by { ext1, rw [to_span_singleton_apply, smul_apply, to_span_singleton_apply, smul_comm], } + +/-- A special case of `to_span_singleton_smul'` for when `R` is commutative. -/ +lemma to_span_singleton_smul (R) {M₁} [comm_semiring R] [add_comm_monoid M₁] [module R M₁] + [topological_space R] [topological_space M₁] [has_continuous_smul R M₁] (c : R) (x : M₁) : + to_span_singleton R (c • x) = c • to_span_singleton R x := +to_span_singleton_smul' R c x + +end to_span_singleton + end semiring section pi From 554bb38de8ded0dafe93b7f18f0bfee6ef77dc5d Mon Sep 17 00:00:00 2001 From: Eric Rodriguez Date: Wed, 21 Jun 2023 22:24:16 +0000 Subject: [PATCH 1213/1291] feat(order/monotone): add `decidable` instances (#19175) These can be written in a very nice, elegant way. Co-authored-by: Yury G. Kudryashov --- .../ascending_descending_sequences.lean | 1 - src/order/monotone/basic.lean | 15 +++++++++++++++ 2 files changed, 15 insertions(+), 1 deletion(-) diff --git a/archive/wiedijk_100_theorems/ascending_descending_sequences.lean b/archive/wiedijk_100_theorems/ascending_descending_sequences.lean index 2c2d3931d8d69..a0d947a8089a5 100644 --- a/archive/wiedijk_100_theorems/ascending_descending_sequences.lean +++ b/archive/wiedijk_100_theorems/ascending_descending_sequences.lean @@ -24,7 +24,6 @@ sequences, increasing, decreasing, Ramsey, Erdos-Szekeres, Erdős–Szekeres, Er variables {α : Type*} [linear_order α] {β : Type*} open function finset -open_locale classical namespace theorems_100 diff --git a/src/order/monotone/basic.lean b/src/order/monotone/basic.lean index 84433c7deff17..9ca070071a7ca 100644 --- a/src/order/monotone/basic.lean +++ b/src/order/monotone/basic.lean @@ -102,6 +102,21 @@ def strict_anti_on (f : α → β) (s : set α) : Prop := end monotone_def +section decidable + +variables [preorder α] [preorder β] {f : α → β} {s : set α} + +instance [i : decidable (∀ a b, a ≤ b → f a ≤ f b)] : decidable (monotone f) := i +instance [i : decidable (∀ a b, a ≤ b → f b ≤ f a)] : decidable (antitone f) := i +instance [i : decidable (∀ a b ∈ s, a ≤ b → f a ≤ f b)] : decidable (monotone_on f s) := i +instance [i : decidable (∀ a b ∈ s, a ≤ b → f b ≤ f a)] : decidable (antitone_on f s) := i +instance [i : decidable (∀ a b, a < b → f a < f b)] : decidable (strict_mono f) := i +instance [i : decidable (∀ a b, a < b → f b < f a)] : decidable (strict_anti f) := i +instance [i : decidable (∀ a b ∈ s, a < b → f a < f b)] : decidable (strict_mono_on f s) := i +instance [i : decidable (∀ a b ∈ s, a < b → f b < f a)] : decidable (strict_anti_on f s) := i + +end decidable + /-! ### Monotonicity on the dual order Strictly, many of the `*_on.dual` lemmas in this section should use `of_dual ⁻¹' s` instead of `s`, From f40476639bac089693a489c9e354ebd75dc0f886 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Thu, 22 Jun 2023 02:23:51 +0000 Subject: [PATCH 1214/1291] feat(analysis/normed_space/multilinear): add lemmas/defs (#19114) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit * Add `simps` here and there. * Add `continuous_multilinear_map.norm_of_subsingleton_le` and `continuous_multilinear_map.nnnorm_of_subsingleton_le`. * Add `continuous_multilinear_map.cod_restrict`. * Add `continuous_multilinear_map.restrict_scalarsₗᵢ`, a `linear_isometry_equiv` version of `continuous_multilinear_map.restrict_scalars`. * Split `continuous_multilinear_map.dom_dom_congr` into 3 definitions: a map, an equivalence, and a linear isometry (the old def). --- src/analysis/normed_space/multilinear.lean | 65 +++++++++++--------- src/topology/algebra/module/multilinear.lean | 25 ++++++++ 2 files changed, 62 insertions(+), 28 deletions(-) diff --git a/src/analysis/normed_space/multilinear.lean b/src/analysis/normed_space/multilinear.lean index da305f925b08a..e2db42edbc5b3 100644 --- a/src/analysis/normed_space/multilinear.lean +++ b/src/analysis/normed_space/multilinear.lean @@ -446,18 +446,23 @@ end section variables (𝕜 G) +lemma norm_of_subsingleton_le [subsingleton ι] (i' : ι) : ‖of_subsingleton 𝕜 G i'‖ ≤ 1 := +op_norm_le_bound _ zero_le_one $ λ m, + by rw [fintype.prod_subsingleton _ i', one_mul, of_subsingleton_apply] + @[simp] lemma norm_of_subsingleton [subsingleton ι] [nontrivial G] (i' : ι) : ‖of_subsingleton 𝕜 G i'‖ = 1 := begin - apply le_antisymm, - { refine op_norm_le_bound _ zero_le_one (λ m, _), - rw [fintype.prod_subsingleton _ i', one_mul, of_subsingleton_apply] }, - { obtain ⟨g, hg⟩ := exists_ne (0 : G), - rw ←norm_ne_zero_iff at hg, - have := (of_subsingleton 𝕜 G i').ratio_le_op_norm (λ _, g), - rwa [fintype.prod_subsingleton _ i', of_subsingleton_apply, div_self hg] at this }, + apply le_antisymm (norm_of_subsingleton_le 𝕜 G i'), + obtain ⟨g, hg⟩ := exists_ne (0 : G), + rw ←norm_ne_zero_iff at hg, + have := (of_subsingleton 𝕜 G i').ratio_le_op_norm (λ _, g), + rwa [fintype.prod_subsingleton _ i', of_subsingleton_apply, div_self hg] at this, end +lemma nnnorm_of_subsingleton_le [subsingleton ι] (i' : ι) : ‖of_subsingleton 𝕜 G i'‖₊ ≤ 1 := +norm_of_subsingleton_le _ _ _ + @[simp] lemma nnnorm_of_subsingleton [subsingleton ι] [nontrivial G] (i' : ι) : ‖of_subsingleton 𝕜 G i'‖₊ = 1 := nnreal.eq $ norm_of_subsingleton _ _ _ @@ -518,18 +523,22 @@ variables {𝕜' : Type*} [nontrivially_normed_field 𝕜'] [normed_algebra 𝕜 variables [normed_space 𝕜' G] [is_scalar_tower 𝕜' 𝕜 G] variables [Π i, normed_space 𝕜' (E i)] [∀ i, is_scalar_tower 𝕜' 𝕜 (E i)] -@[simp] lemma norm_restrict_scalars : ‖f.restrict_scalars 𝕜'‖ = ‖f‖ := -by simp only [norm_def, coe_restrict_scalars] +@[simp] lemma norm_restrict_scalars : ‖f.restrict_scalars 𝕜'‖ = ‖f‖ := rfl variable (𝕜') -/-- `continuous_multilinear_map.restrict_scalars` as a `continuous_multilinear_map`. -/ -def restrict_scalars_linear : - continuous_multilinear_map 𝕜 E G →L[𝕜'] continuous_multilinear_map 𝕜' E G := -linear_map.mk_continuous +/-- `continuous_multilinear_map.restrict_scalars` as a `linear_isometry`. -/ +def restrict_scalarsₗᵢ : + continuous_multilinear_map 𝕜 E G →ₗᵢ[𝕜'] continuous_multilinear_map 𝕜' E G := { to_fun := restrict_scalars 𝕜', map_add' := λ m₁ m₂, rfl, - map_smul' := λ c m, rfl } 1 $ λ f, by simp + map_smul' := λ c m, rfl, + norm_map' := λ f, rfl } + +/-- `continuous_multilinear_map.restrict_scalars` as a `continuous_linear_map`. -/ +def restrict_scalars_linear : + continuous_multilinear_map 𝕜 E G →L[𝕜'] continuous_multilinear_map 𝕜' E G := +(restrict_scalarsₗᵢ 𝕜').to_continuous_linear_map variable {𝕜'} @@ -922,6 +931,7 @@ rfl /-- Flip arguments in `f : G →L[𝕜] continuous_multilinear_map 𝕜 E G'` to get `continuous_multilinear_map 𝕜 E (G →L[𝕜] G')` -/ +@[simps apply_apply] def flip_multilinear (f : G →L[𝕜] continuous_multilinear_map 𝕜 E G') : continuous_multilinear_map 𝕜 E (G →L[𝕜] G') := multilinear_map.mk_continuous @@ -1503,23 +1513,22 @@ namespace continuous_multilinear_map variables (𝕜 G G') +-- fails to unify without `@`; TODO: try again in Lean 4 +@[simp] theorem norm_dom_dom_congr (σ : ι ≃ ι') (f : continuous_multilinear_map 𝕜 (λ _ : ι, G) G') : + ‖@dom_dom_congr 𝕜 ι G G' _ _ _ _ _ _ _ ι' σ f‖ = ‖f‖ := +by simp only [norm_def, linear_equiv.coe_mk, ← σ.prod_comp, + (σ.arrow_congr (equiv.refl G)).surjective.forall, dom_dom_congr_apply, equiv.arrow_congr_apply, + equiv.coe_refl, comp.left_id, comp_app, equiv.symm_apply_apply, id] + /-- An equivalence of the index set defines a linear isometric equivalence between the spaces of multilinear maps. -/ -def dom_dom_congr (σ : ι ≃ ι') : +def dom_dom_congrₗᵢ (σ : ι ≃ ι') : continuous_multilinear_map 𝕜 (λ _ : ι, G) G' ≃ₗᵢ[𝕜] continuous_multilinear_map 𝕜 (λ _ : ι', G) G' := -linear_isometry_equiv.of_bounds - { to_fun := λ f, (multilinear_map.dom_dom_congr σ f.to_multilinear_map).mk_continuous ‖f‖ $ - λ m, (f.le_op_norm (λ i, m (σ i))).trans_eq $ by rw [← σ.prod_comp], - inv_fun := λ f, (multilinear_map.dom_dom_congr σ.symm f.to_multilinear_map).mk_continuous ‖f‖ $ - λ m, (f.le_op_norm (λ i, m (σ.symm i))).trans_eq $ by rw [← σ.symm.prod_comp], - left_inv := λ f, ext $ λ m, congr_arg f $ by simp only [σ.symm_apply_apply], - right_inv := λ f, ext $ λ m, congr_arg f $ by simp only [σ.apply_symm_apply], - map_add' := λ f g, rfl, - map_smul' := λ c f, rfl } - (λ f, multilinear_map.mk_continuous_norm_le _ (norm_nonneg f) _) - (λ f, multilinear_map.mk_continuous_norm_le _ (norm_nonneg f) _) - + { map_add' := λ _ _, rfl, + map_smul' := λ _ _, rfl, + norm_map' := norm_dom_dom_congr 𝕜 G G' σ, + .. dom_dom_congr_equiv σ } variables {𝕜 G G'} section @@ -1589,7 +1598,7 @@ values in the space of continuous multilinear maps of `l` variables. -/ def curry_fin_finset {k l n : ℕ} {s : finset (fin n)} (hk : s.card = k) (hl : sᶜ.card = l) : (G [×n]→L[𝕜] G') ≃ₗᵢ[𝕜] (G [×k]→L[𝕜] G [×l]→L[𝕜] G') := -(dom_dom_congr 𝕜 G G' (fin_sum_equiv_of_finset hk hl).symm).trans +(dom_dom_congrₗᵢ 𝕜 G G' (fin_sum_equiv_of_finset hk hl).symm).trans (curry_sum_equiv 𝕜 (fin k) (fin l) G G') variables {𝕜 G G'} diff --git a/src/topology/algebra/module/multilinear.lean b/src/topology/algebra/module/multilinear.lean index efa86ed869ffc..ef1fead7e6fc3 100644 --- a/src/topology/algebra/module/multilinear.lean +++ b/src/topology/algebra/module/multilinear.lean @@ -217,6 +217,12 @@ lemma pi_apply {ι' : Type*} {M' : ι' → Type*} [Π i, add_comm_monoid (M' i)] pi f m j = f j m := rfl +/-- Restrict the codomain of a continuous multilinear map to a submodule. -/ +@[simps to_multilinear_map apply_coe] +def cod_restrict (f : continuous_multilinear_map R M₁ M₂) (p : submodule R M₂) (h : ∀ v, f v ∈ p) : + continuous_multilinear_map R M₁ p := +⟨f.1.cod_restrict p h, f.cont.subtype_mk _⟩ + section variables (R M₂) @@ -276,6 +282,25 @@ def pi_equiv {ι' : Type*} {M' : ι' → Type*} [Π i, add_comm_monoid (M' i)] left_inv := λ f, by { ext, refl }, right_inv := λ f, by { ext, refl } } +/-- An equivalence of the index set defines an equivalence between the spaces of continuous +multilinear maps. This is the forward map of this equivalence. -/ +@[simps to_multilinear_map apply] +def dom_dom_congr {ι' : Type*} (e : ι ≃ ι') (f : continuous_multilinear_map R (λ _ : ι, M₂) M₃) : + continuous_multilinear_map R (λ _ : ι', M₂) M₃ := +{ to_multilinear_map := f.dom_dom_congr e, + cont := f.cont.comp $ continuous_pi $ λ _, continuous_apply _ } + +/-- An equivalence of the index set defines an equivalence between the spaces of continuous +multilinear maps. In case of normed spaces, this is a linear isometric equivalence, see +`continuous.multilinear_map.dom_dom_congrₗᵢ`. -/ +@[simps] +def dom_dom_congr_equiv {ι' : Type*} (e : ι ≃ ι') : + continuous_multilinear_map R (λ _ : ι, M₂) M₃ ≃ continuous_multilinear_map R (λ _ : ι', M₂) M₃ := +{ to_fun := dom_dom_congr e, + inv_fun := dom_dom_congr e.symm, + left_inv := λ _, ext $ λ _, by simp, + right_inv := λ _, ext $ λ _, by simp } + /-- In the specific case of continuous multilinear maps on spaces indexed by `fin (n+1)`, where one can build an element of `Π(i : fin (n+1)), M i` using `cons`, one can express directly the additivity of a multilinear map along the first variable. -/ From 1b089e3bdc3ce6b39cd472543474a0a137128c6c Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Thu, 22 Jun 2023 05:03:22 +0000 Subject: [PATCH 1215/1291] chore(*): add mathlib4 synchronization comments (#19211) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.category.Group.adjunctions` * `analysis.constant_speed` * `analysis.convolution` * `category_theory.category.PartialFun` * `control.bitraversable.instances` * `control.uliftable` * `data.nat.nth` * `field_theory.cardinality` * `field_theory.finite.galois_field` * `field_theory.finite.trace` * `number_theory.function_field` * `number_theory.number_field.basic` * `number_theory.number_field.norm` * `number_theory.prime_counting` * `number_theory.zeta_values` * `probability.ident_distrib` * `probability.kernel.cond_distrib` * `probability.kernel.condexp` * `probability.moments` * `probability.strong_law` * `probability.variance` * `representation_theory.Rep` * `ring_theory.complex` * `ring_theory.valuation.ramification_group` * `ring_theory.valuation.valuation_subring` * `set_theory.ordinal.notation` --- src/algebra/category/Group/adjunctions.lean | 3 +++ src/analysis/constant_speed.lean | 3 +++ src/analysis/convolution.lean | 3 +++ src/category_theory/category/PartialFun.lean | 3 +++ src/control/bitraversable/instances.lean | 3 +++ src/control/uliftable.lean | 3 +++ src/data/nat/nth.lean | 3 +++ src/field_theory/cardinality.lean | 3 +++ src/field_theory/finite/galois_field.lean | 3 +++ src/field_theory/finite/trace.lean | 3 +++ src/number_theory/function_field.lean | 3 +++ src/number_theory/number_field/basic.lean | 3 +++ src/number_theory/number_field/norm.lean | 3 +++ src/number_theory/prime_counting.lean | 3 +++ src/number_theory/zeta_values.lean | 3 +++ src/probability/ident_distrib.lean | 3 +++ src/probability/kernel/cond_distrib.lean | 3 +++ src/probability/kernel/condexp.lean | 3 +++ src/probability/moments.lean | 3 +++ src/probability/strong_law.lean | 3 +++ src/probability/variance.lean | 3 +++ src/representation_theory/Rep.lean | 3 +++ src/ring_theory/complex.lean | 5 ++++- src/ring_theory/valuation/ramification_group.lean | 3 +++ src/ring_theory/valuation/valuation_subring.lean | 3 +++ src/set_theory/ordinal/notation.lean | 3 +++ 26 files changed, 79 insertions(+), 1 deletion(-) diff --git a/src/algebra/category/Group/adjunctions.lean b/src/algebra/category/Group/adjunctions.lean index 64f2caf41d79c..8aa601a197357 100644 --- a/src/algebra/category/Group/adjunctions.lean +++ b/src/algebra/category/Group/adjunctions.lean @@ -9,6 +9,9 @@ import group_theory.free_abelian_group /-! # Adjunctions regarding the category of (abelian) groups +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains construction of basic adjunctions concerning the category of groups and the category of abelian groups. diff --git a/src/analysis/constant_speed.lean b/src/analysis/constant_speed.lean index c6ac9e016baf7..aa5b84190414f 100644 --- a/src/analysis/constant_speed.lean +++ b/src/analysis/constant_speed.lean @@ -9,6 +9,9 @@ import tactic.swap_var /-! # Constant speed +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the notion of constant (and unit) speed for a function `f : ℝ → E` with pseudo-emetric structure on `E` with respect to a set `s : set ℝ` and "speed" `l : ℝ≥0`, and shows that if `f` has locally bounded variation on `s`, it can be obtained (up to distance zero, on `s`), diff --git a/src/analysis/convolution.lean b/src/analysis/convolution.lean index edcbfe3a605bb..74b30a8db71c0 100644 --- a/src/analysis/convolution.lean +++ b/src/analysis/convolution.lean @@ -14,6 +14,9 @@ import measure_theory.integral.interval_integral /-! # Convolution of functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the convolution on two functions, i.e. `x ↦ ∫ f(t)g(x - t) ∂t`. In the general case, these functions can be vector-valued, and have an arbitrary (additive) group as domain. We use a continuous bilinear operation `L` on these function values as diff --git a/src/category_theory/category/PartialFun.lean b/src/category_theory/category/PartialFun.lean index bef9af2a87470..46ffcb1b33c11 100644 --- a/src/category_theory/category/PartialFun.lean +++ b/src/category_theory/category/PartialFun.lean @@ -9,6 +9,9 @@ import data.pfun /-! # The category of types with partial functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This defines `PartialFun`, the category of types equipped with partial functions. This category is classically equivalent to the category of pointed types. The reason it doesn't hold diff --git a/src/control/bitraversable/instances.lean b/src/control/bitraversable/instances.lean index 851bc71c60b8a..1b6bb4da9175d 100644 --- a/src/control/bitraversable/instances.lean +++ b/src/control/bitraversable/instances.lean @@ -9,6 +9,9 @@ import control.traversable.lemmas /-! # Bitraversable instances +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file provides `bitraversable` instances for concrete bifunctors: * `prod` * `sum` diff --git a/src/control/uliftable.lean b/src/control/uliftable.lean index 3cdadee22171d..104ed8c3b5399 100644 --- a/src/control/uliftable.lean +++ b/src/control/uliftable.lean @@ -12,6 +12,9 @@ import tactic.interactive /-! # Universe lifting for type families +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Some functors such as `option` and `list` are universe polymorphic. Unlike type polymorphism where `option α` is a function application and reasoning and generalizations that apply to functions can be used, `option.{u}` and `option.{v}` diff --git a/src/data/nat/nth.lean b/src/data/nat/nth.lean index c7802520e6ce0..42541ced019d3 100644 --- a/src/data/nat/nth.lean +++ b/src/data/nat/nth.lean @@ -10,6 +10,9 @@ import order.order_iso_nat /-! # The `n`th Number Satisfying a Predicate +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines a function for "what is the `n`th number that satisifies a given predicate `p`", and provides lemmas that deal with this function and its connection to `nat.count`. diff --git a/src/field_theory/cardinality.lean b/src/field_theory/cardinality.lean index 17a1c271f3e9b..e01d6b48ab64b 100644 --- a/src/field_theory/cardinality.lean +++ b/src/field_theory/cardinality.lean @@ -15,6 +15,9 @@ import set_theory.cardinal.divisibility /-! # Cardinality of Fields +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we show all the possible cardinalities of fields. All infinite cardinals can harbour a field structure, and so can all types with prime power cardinalities, and this is sharp. diff --git a/src/field_theory/finite/galois_field.lean b/src/field_theory/finite/galois_field.lean index 8737cefa46b92..d787c29c4b3c7 100644 --- a/src/field_theory/finite/galois_field.lean +++ b/src/field_theory/finite/galois_field.lean @@ -13,6 +13,9 @@ import field_theory.splitting_field.is_splitting_field /-! # Galois fields +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + If `p` is a prime number, and `n` a natural number, then `galois_field p n` is defined as the splitting field of `X^(p^n) - X` over `zmod p`. It is a finite field with `p ^ n` elements. diff --git a/src/field_theory/finite/trace.lean b/src/field_theory/finite/trace.lean index 42c334451a5f2..59dd268d8dc49 100644 --- a/src/field_theory/finite/trace.lean +++ b/src/field_theory/finite/trace.lean @@ -9,6 +9,9 @@ import field_theory.finite.galois_field /-! # The trace map for finite fields +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We state the fact that the trace map from a finite field of characteristic `p` to `zmod p` is nondegenerate. diff --git a/src/number_theory/function_field.lean b/src/number_theory/function_field.lean index 74b7f229c433a..491ce0bc15592 100644 --- a/src/number_theory/function_field.lean +++ b/src/number_theory/function_field.lean @@ -12,6 +12,9 @@ import topology.algebra.valued_field /-! # Function fields +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines a function field and the ring of integers corresponding to it. ## Main definitions diff --git a/src/number_theory/number_field/basic.lean b/src/number_theory/number_field/basic.lean index 537628af9ebf9..a426ce89fb1e7 100644 --- a/src/number_theory/number_field/basic.lean +++ b/src/number_theory/number_field/basic.lean @@ -8,6 +8,9 @@ import ring_theory.dedekind_domain.integral_closure /-! # Number fields + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. This file defines a number field and the ring of integers corresponding to it. ## Main definitions diff --git a/src/number_theory/number_field/norm.lean b/src/number_theory/number_field/norm.lean index 8476cfc5dff86..1699edeb6c4c3 100644 --- a/src/number_theory/number_field/norm.lean +++ b/src/number_theory/number_field/norm.lean @@ -9,6 +9,9 @@ import ring_theory.norm /-! # Norm in number fields + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. Given a finite extension of number fields, we define the norm morphism as a function between the rings of integers. diff --git a/src/number_theory/prime_counting.lean b/src/number_theory/prime_counting.lean index d811b98cd28f0..6a699a1ae127d 100644 --- a/src/number_theory/prime_counting.lean +++ b/src/number_theory/prime_counting.lean @@ -13,6 +13,9 @@ import data.nat.nth /-! # The Prime Counting Function +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define the prime counting function: the function on natural numbers that returns the number of primes less than or equal to its input. diff --git a/src/number_theory/zeta_values.lean b/src/number_theory/zeta_values.lean index bb90b628f09d4..296f70123d559 100644 --- a/src/number_theory/zeta_values.lean +++ b/src/number_theory/zeta_values.lean @@ -12,6 +12,9 @@ import analysis.p_series /-! # Critical values of the Riemann zeta function +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove formulae for the critical values of `ζ(s)`, and more generally of Hurwitz zeta functions, in terms of Bernoulli polynomials. diff --git a/src/probability/ident_distrib.lean b/src/probability/ident_distrib.lean index 75775ebd5c98a..022eafdb4b26b 100644 --- a/src/probability/ident_distrib.lean +++ b/src/probability/ident_distrib.lean @@ -9,6 +9,9 @@ import measure_theory.function.uniform_integrable /-! # Identically distributed random variables +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Two random variables defined on two (possibly different) probability spaces but taking value in the same space are *identically distributed* if their distributions (i.e., the image probability measures on the target space) coincide. We define this concept and establish its basic properties diff --git a/src/probability/kernel/cond_distrib.lean b/src/probability/kernel/cond_distrib.lean index 6da8554ccaf4b..d1a141b01b8e4 100644 --- a/src/probability/kernel/cond_distrib.lean +++ b/src/probability/kernel/cond_distrib.lean @@ -9,6 +9,9 @@ import probability.notation /-! # Regular conditional probability distribution +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define the regular conditional probability distribution of `Y : α → Ω` given `X : α → β`, where `Ω` is a standard Borel space. This is a `kernel β Ω` such that for almost all `a`, `cond_distrib` evaluated at `X a` and a measurable set `s` is equal to the conditional expectation diff --git a/src/probability/kernel/condexp.lean b/src/probability/kernel/condexp.lean index d860a20e4c3ac..d63c7bb03b265 100644 --- a/src/probability/kernel/condexp.lean +++ b/src/probability/kernel/condexp.lean @@ -8,6 +8,9 @@ import probability.kernel.cond_distrib /-! # Kernel associated with a conditional expectation +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define `condexp_kernel μ m`, a kernel from `Ω` to `Ω` such that for all integrable functions `f`, `μ[f | m] =ᵐ[μ] λ ω, ∫ y, f y ∂(condexp_kernel μ m ω)`. diff --git a/src/probability/moments.lean b/src/probability/moments.lean index 392a55356ec10..64126645bbb0a 100644 --- a/src/probability/moments.lean +++ b/src/probability/moments.lean @@ -9,6 +9,9 @@ import probability.variance /-! # Moments and moment generating function +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions * `probability_theory.moment X p μ`: `p`th moment of a real random variable `X` with respect to diff --git a/src/probability/strong_law.lean b/src/probability/strong_law.lean index e09fa4393ca78..53020a09142fa 100644 --- a/src/probability/strong_law.lean +++ b/src/probability/strong_law.lean @@ -13,6 +13,9 @@ import analysis.asymptotics.specific_asymptotics /-! # The strong law of large numbers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We prove the strong law of large numbers, in `probability_theory.strong_law_ae`: If `X n` is a sequence of independent identically distributed integrable real-valued random variables, then `∑ i in range n, X i / n` converges almost surely to `𝔼[X 0]`. diff --git a/src/probability/variance.lean b/src/probability/variance.lean index a0b0c08c87d91..bf41ae83ad847 100644 --- a/src/probability/variance.lean +++ b/src/probability/variance.lean @@ -10,6 +10,9 @@ import measure_theory.function.l2_space /-! # Variance of random variables +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define the variance of a real-valued random variable as `Var[X] = 𝔼[(X - 𝔼[X])^2]` (in the `probability_theory` locale). diff --git a/src/representation_theory/Rep.lean b/src/representation_theory/Rep.lean index 2f801e5652d25..cd4ef0178d7bd 100644 --- a/src/representation_theory/Rep.lean +++ b/src/representation_theory/Rep.lean @@ -14,6 +14,9 @@ import category_theory.closed.functor_category /-! # `Rep k G` is the category of `k`-linear representations of `G`. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + If `V : Rep k G`, there is a coercion that allows you to treat `V` as a type, and this type comes equipped with a `module k V` instance. Also `V.ρ` gives the homomorphism `G →* (V →ₗ[k] V)`. diff --git a/src/ring_theory/complex.lean b/src/ring_theory/complex.lean index 598d189f045c2..88b939cc3fb71 100644 --- a/src/ring_theory/complex.lean +++ b/src/ring_theory/complex.lean @@ -7,7 +7,10 @@ import data.complex.module import ring_theory.norm import ring_theory.trace -/-! # Lemmas about `algebra.trace` and `algebra.norm` on `ℂ` -/ +/-! # Lemmas about `algebra.trace` and `algebra.norm` on `ℂ` + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4.-/ open complex diff --git a/src/ring_theory/valuation/ramification_group.lean b/src/ring_theory/valuation/ramification_group.lean index 9276b19d54fb1..ed286cb001bc3 100644 --- a/src/ring_theory/valuation/ramification_group.lean +++ b/src/ring_theory/valuation/ramification_group.lean @@ -9,6 +9,9 @@ import ring_theory.valuation.valuation_subring /-! # Ramification groups +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The decomposition subgroup and inertia subgroups. TODO: Define higher ramification groups in lower numbering diff --git a/src/ring_theory/valuation/valuation_subring.lean b/src/ring_theory/valuation/valuation_subring.lean index faf017aa2555c..156e1ae708a0e 100644 --- a/src/ring_theory/valuation/valuation_subring.lean +++ b/src/ring_theory/valuation/valuation_subring.lean @@ -12,6 +12,9 @@ import algebraic_geometry.prime_spectrum.basic # Valuation subrings of a field +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Projects The order structure on `valuation_subring K`. diff --git a/src/set_theory/ordinal/notation.lean b/src/set_theory/ordinal/notation.lean index cb4d422862155..606e136c196b7 100644 --- a/src/set_theory/ordinal/notation.lean +++ b/src/set_theory/ordinal/notation.lean @@ -9,6 +9,9 @@ import set_theory.ordinal.principal /-! # Ordinal notation +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Constructive ordinal arithmetic for ordinals below `ε₀`. We define a type `onote`, with constructors `0 : onote` and `onote.oadd e n a` representing From d30d31261cdb4d2f5e612eabc3c4bf45556350d5 Mon Sep 17 00:00:00 2001 From: Eric Rodriguez Date: Thu, 22 Jun 2023 12:28:45 +0000 Subject: [PATCH 1216/1291] feat(group_theory): simple lemmas for Wedderburn (#18862) These lemmas are a bit disparate, but they are all useful for Wedderburn's little theorem. Co-authored-by: Eric Rodriguez <37984851+ericrbg@users.noreply.github.com> --- src/group_theory/group_action/basic.lean | 3 +++ src/group_theory/group_action/conj_act.lean | 6 ++++++ src/group_theory/subgroup/basic.lean | 10 ++++++++++ 3 files changed, 19 insertions(+) diff --git a/src/group_theory/group_action/basic.lean b/src/group_theory/group_action/basic.lean index 767fc4042cae6..c9dfe90ea91b9 100644 --- a/src/group_theory/group_action/basic.lean +++ b/src/group_theory/group_action/basic.lean @@ -192,6 +192,9 @@ local attribute [instance] orbit_rel variables {α} {β} +@[to_additive] +lemma orbit_rel_apply {x y : β} : (orbit_rel α β).rel x y ↔ x ∈ orbit α y := iff.rfl + /-- When you take a set `U` in `β`, push it down to the quotient, and pull back, you get the union of the orbit of `U` under `α`. -/ @[to_additive "When you take a set `U` in `β`, push it down to the quotient, and pull back, you get diff --git a/src/group_theory/group_action/conj_act.lean b/src/group_theory/group_action/conj_act.lean index 94d3aa0d6bdf4..12809e56e22d4 100644 --- a/src/group_theory/group_action/conj_act.lean +++ b/src/group_theory/group_action/conj_act.lean @@ -187,6 +187,12 @@ begin simp [mem_center_iff, smul_def, mul_inv_eq_iff_eq_mul] end +@[simp] lemma mem_orbit_conj_act {g h : G} : g ∈ orbit (conj_act G) h ↔ is_conj g h := +by { rw [is_conj_comm, is_conj_iff, mem_orbit_iff], refl } + +lemma orbit_rel_conj_act : (orbit_rel (conj_act G) G).rel = is_conj := +funext₂ $ λ g h, by rw [orbit_rel_apply, mem_orbit_conj_act] + lemma stabilizer_eq_centralizer (g : G) : stabilizer (conj_act G) g = (zpowers g).centralizer := le_antisymm (le_centralizer_iff.mp (zpowers_le.mpr (λ x, mul_inv_eq_iff_eq_mul.mp))) (λ x h, mul_inv_eq_of_eq_mul (h g (mem_zpowers g)).symm) diff --git a/src/group_theory/subgroup/basic.lean b/src/group_theory/subgroup/basic.lean index 42878c5673ac4..464a9504a5aaa 100644 --- a/src/group_theory/subgroup/basic.lean +++ b/src/group_theory/subgroup/basic.lean @@ -2676,6 +2676,16 @@ begin exact subset_normal_closure (set.mem_singleton _), end +variables {M : Type*} [monoid M] + +lemma eq_of_left_mem_center {g h : M} (H : is_conj g h) (Hg : g ∈ set.center M) : + g = h := +by { rcases H with ⟨u, hu⟩, rwa [← u.mul_left_inj, ← Hg u], } + +lemma eq_of_right_mem_center {g h : M} (H : is_conj g h) (Hh : h ∈ set.center M) : + g = h := +(H.symm.eq_of_left_mem_center Hh).symm + end is_conj assert_not_exists multiset From 57f9349f2fe19d2de7207e99b0341808d977cdcf Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Thu, 22 Jun 2023 23:00:27 +0000 Subject: [PATCH 1217/1291] chore(analysis/complex/upper_half_plane,number_theory/modular_forms): reduce dependency on manifolds (#19212) Move some lemmas that use `mfderiv` or `mdifferentiable` to new files. --- .../complex/upper_half_plane/manifold.lean | 32 +++++++++++++++++++ .../complex/upper_half_plane/topology.lean | 17 +--------- src/number_theory/modular_forms/basic.lean | 3 +- .../basic.lean} | 10 ++---- .../modular_forms/jacobi_theta/manifold.lean | 23 +++++++++++++ src/number_theory/zeta_function.lean | 2 +- 6 files changed, 61 insertions(+), 26 deletions(-) create mode 100644 src/analysis/complex/upper_half_plane/manifold.lean rename src/number_theory/modular_forms/{jacobi_theta.lean => jacobi_theta/basic.lean} (95%) create mode 100644 src/number_theory/modular_forms/jacobi_theta/manifold.lean diff --git a/src/analysis/complex/upper_half_plane/manifold.lean b/src/analysis/complex/upper_half_plane/manifold.lean new file mode 100644 index 0000000000000..b5a8b7ec6de4b --- /dev/null +++ b/src/analysis/complex/upper_half_plane/manifold.lean @@ -0,0 +1,32 @@ +/- +Copyright (c) 2022 Chris Birkbeck. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Chris Birkbeck +-/ +import analysis.complex.upper_half_plane.topology +import geometry.manifold.cont_mdiff_mfderiv +/-! +# Manifold structure on the upper half plane. + +In this file we define the complex manifold structure on the upper half-plane. +-/ + +open_locale upper_half_plane manifold + +namespace upper_half_plane + +noncomputable instance : charted_space ℂ ℍ := +upper_half_plane.open_embedding_coe.singleton_charted_space + +instance : smooth_manifold_with_corners 𝓘(ℂ) ℍ := +upper_half_plane.open_embedding_coe.singleton_smooth_manifold_with_corners 𝓘(ℂ) + +/-- The inclusion map `ℍ → ℂ` is a smooth map of manifolds. -/ +lemma smooth_coe : smooth 𝓘(ℂ) 𝓘(ℂ) (coe : ℍ → ℂ) := +λ x, cont_mdiff_at_ext_chart_at + +/-- The inclusion map `ℍ → ℂ` is a differentiable map of manifolds. -/ +lemma mdifferentiable_coe : mdifferentiable 𝓘(ℂ) 𝓘(ℂ) (coe : ℍ → ℂ) := +smooth_coe.mdifferentiable + +end upper_half_plane diff --git a/src/analysis/complex/upper_half_plane/topology.lean b/src/analysis/complex/upper_half_plane/topology.lean index c441e99fcfe93..418466831de71 100644 --- a/src/analysis/complex/upper_half_plane/topology.lean +++ b/src/analysis/complex/upper_half_plane/topology.lean @@ -9,7 +9,6 @@ import analysis.convex.normed import analysis.convex.complex import analysis.complex.re_im_topology import topology.homotopy.contractible -import geometry.manifold.cont_mdiff_mfderiv /-! # Topology on the upper half plane @@ -20,7 +19,7 @@ various instances. noncomputable theory open set filter function topological_space complex -open_locale filter topology upper_half_plane manifold +open_locale filter topology upper_half_plane namespace upper_half_plane @@ -58,18 +57,4 @@ end instance : locally_compact_space ℍ := open_embedding_coe.locally_compact_space -instance upper_half_plane.charted_space : charted_space ℂ ℍ := -upper_half_plane.open_embedding_coe.singleton_charted_space - -instance upper_half_plane.smooth_manifold_with_corners : smooth_manifold_with_corners 𝓘(ℂ) ℍ := -upper_half_plane.open_embedding_coe.singleton_smooth_manifold_with_corners 𝓘(ℂ) - -/-- The inclusion map `ℍ → ℂ` is a smooth map of manifolds. -/ -lemma smooth_coe : smooth 𝓘(ℂ) 𝓘(ℂ) (coe : ℍ → ℂ) := -λ x, cont_mdiff_at_ext_chart_at - -/-- The inclusion map `ℍ → ℂ` is a differentiable map of manifolds. -/ -lemma mdifferentiable_coe : mdifferentiable 𝓘(ℂ) 𝓘(ℂ) (coe : ℍ → ℂ) := -smooth_coe.mdifferentiable - end upper_half_plane diff --git a/src/number_theory/modular_forms/basic.lean b/src/number_theory/modular_forms/basic.lean index 693f1d5696ef8..1d1346c5dfee0 100644 --- a/src/number_theory/modular_forms/basic.lean +++ b/src/number_theory/modular_forms/basic.lean @@ -3,9 +3,8 @@ Copyright (c) 2022 Chris Birkbeck. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Birkbeck -/ - import analysis.complex.upper_half_plane.functions_bounded_at_infty -import analysis.complex.upper_half_plane.topology +import analysis.complex.upper_half_plane.manifold import number_theory.modular_forms.slash_invariant_forms /-! # Modular forms diff --git a/src/number_theory/modular_forms/jacobi_theta.lean b/src/number_theory/modular_forms/jacobi_theta/basic.lean similarity index 95% rename from src/number_theory/modular_forms/jacobi_theta.lean rename to src/number_theory/modular_forms/jacobi_theta/basic.lean index 1afe11de17ba4..5dfc5e66f6a90 100644 --- a/src/number_theory/modular_forms/jacobi_theta.lean +++ b/src/number_theory/modular_forms/jacobi_theta/basic.lean @@ -19,9 +19,8 @@ and proves the modular transformation properties `θ (τ + 2) = θ τ` and show that `θ` is differentiable on `ℍ`, and `θ(τ) - 1` has exponential decay as `im τ → ∞`. -/ -open complex real asymptotics - -open_locale real big_operators upper_half_plane manifold +open complex real asymptotics filter +open_locale real big_operators upper_half_plane /-- Jacobi's theta function `∑' (n : ℤ), exp (π * I * n ^ 2 * τ)`. -/ noncomputable def jacobi_theta (z : ℂ) : ℂ := ∑' (n : ℤ), cexp (π * I * n ^ 2 * z) @@ -151,7 +150,7 @@ end /-- The norm of `jacobi_theta τ - 1` decays exponentially as `im τ → ∞`. -/ lemma is_O_at_im_infty_jacobi_theta_sub_one : - is_O (filter.comap im filter.at_top) (λ τ, jacobi_theta τ - 1) (λ τ, rexp (-π * τ.im)) := + (λ τ, jacobi_theta τ - 1) =O[comap im at_top] (λ τ, rexp (-π * τ.im)) := begin simp_rw [is_O, is_O_with, filter.eventually_comap, filter.eventually_at_top], refine ⟨2 / (1 - rexp (-π)), 1, λ y hy z hz, (norm_jacobi_theta_sub_one_le @@ -181,8 +180,5 @@ begin exact differentiable_on_tsum_of_summable_norm bd_s h1 h2 (λ i w hw, le_bd (le_of_lt hw) i), end -lemma mdifferentiable_jacobi_theta : mdifferentiable 𝓘(ℂ) 𝓘(ℂ) (jacobi_theta ∘ coe : ℍ → ℂ) := -λ τ, (differentiable_at_jacobi_theta τ.2).mdifferentiable_at.comp τ τ.mdifferentiable_coe - lemma continuous_at_jacobi_theta {z : ℂ} (hz : 0 < im z) : continuous_at jacobi_theta z := (differentiable_at_jacobi_theta hz).continuous_at diff --git a/src/number_theory/modular_forms/jacobi_theta/manifold.lean b/src/number_theory/modular_forms/jacobi_theta/manifold.lean new file mode 100644 index 0000000000000..67094c9757bca --- /dev/null +++ b/src/number_theory/modular_forms/jacobi_theta/manifold.lean @@ -0,0 +1,23 @@ +/- +Copyright (c) 2023 David Loeffler. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: David Loeffler +-/ +import number_theory.modular_forms.jacobi_theta.basic +import analysis.complex.upper_half_plane.manifold + +/-! +# Manifold differentiability of the Jacobi's theta function + +In this file we reformulate differentiability of the Jacobi's theta function in terms of manifold +differentiability. + +## TODO + +Prove smoothness (in terms of `smooth`). +-/ + +open_locale upper_half_plane manifold + +lemma mdifferentiable_jacobi_theta : mdifferentiable 𝓘(ℂ) 𝓘(ℂ) (jacobi_theta ∘ coe : ℍ → ℂ) := +λ τ, (differentiable_at_jacobi_theta τ.2).mdifferentiable_at.comp τ τ.mdifferentiable_coe diff --git a/src/number_theory/zeta_function.lean b/src/number_theory/zeta_function.lean index 4254a814864f1..b3dc94155d253 100644 --- a/src/number_theory/zeta_function.lean +++ b/src/number_theory/zeta_function.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: David Loeffler -/ import analysis.special_functions.gamma.beta -import number_theory.modular_forms.jacobi_theta +import number_theory.modular_forms.jacobi_theta.basic import number_theory.zeta_values /-! From 5d0c76894ada7940957143163d7b921345474cbc Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Fri, 23 Jun 2023 06:33:52 +0000 Subject: [PATCH 1218/1291] chore(*): add mathlib4 synchronization comments (#19213) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebraic_geometry.Gamma_Spec_adjunction` * `algebraic_geometry.open_immersion.Scheme` * `analysis.calculus.implicit` * `analysis.matrix` * `analysis.normed_space.matrix_exponential` * `analysis.normed_space.star.matrix` * `data.matrix.invertible` * `linear_algebra.clifford_algebra.basic` * `linear_algebra.exterior_algebra.basic` * `number_theory.cyclotomic.basic` * `number_theory.cyclotomic.primitive_roots` * `number_theory.number_field.embeddings` * `number_theory.number_field.units` * `number_theory.zsqrtd.gaussian_int` * `ring_theory.dedekind_domain.S_integer` * `ring_theory.dedekind_domain.adic_valuation` * `ring_theory.dedekind_domain.finite_adele_ring` * `ring_theory.discriminant` * `ring_theory.ideal.norm` --- src/algebraic_geometry/Gamma_Spec_adjunction.lean | 3 +++ src/algebraic_geometry/open_immersion/Scheme.lean | 3 +++ src/analysis/calculus/implicit.lean | 3 +++ src/analysis/matrix.lean | 3 +++ src/analysis/normed_space/matrix_exponential.lean | 3 +++ src/analysis/normed_space/star/matrix.lean | 3 +++ src/data/matrix/invertible.lean | 3 +++ src/linear_algebra/clifford_algebra/basic.lean | 3 +++ src/linear_algebra/exterior_algebra/basic.lean | 3 +++ src/number_theory/cyclotomic/basic.lean | 3 +++ src/number_theory/cyclotomic/primitive_roots.lean | 3 +++ src/number_theory/number_field/embeddings.lean | 3 +++ src/number_theory/number_field/units.lean | 3 +++ src/number_theory/zsqrtd/gaussian_int.lean | 3 +++ src/ring_theory/dedekind_domain/S_integer.lean | 3 +++ src/ring_theory/dedekind_domain/adic_valuation.lean | 3 +++ src/ring_theory/dedekind_domain/finite_adele_ring.lean | 3 +++ src/ring_theory/discriminant.lean | 3 +++ src/ring_theory/ideal/norm.lean | 3 +++ 19 files changed, 57 insertions(+) diff --git a/src/algebraic_geometry/Gamma_Spec_adjunction.lean b/src/algebraic_geometry/Gamma_Spec_adjunction.lean index ebe6c4c29ea3f..566db61904700 100644 --- a/src/algebraic_geometry/Gamma_Spec_adjunction.lean +++ b/src/algebraic_geometry/Gamma_Spec_adjunction.lean @@ -10,6 +10,9 @@ import category_theory.adjunction.reflective /-! # Adjunction between `Γ` and `Spec` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define the adjunction `Γ_Spec.adjunction : Γ ⊣ Spec` by defining the unit (`to_Γ_Spec`, in multiple steps in this file) and counit (done in Spec.lean) and checking that they satisfy the left and right triangle identities. The constructions and proofs make use of diff --git a/src/algebraic_geometry/open_immersion/Scheme.lean b/src/algebraic_geometry/open_immersion/Scheme.lean index 571f52b4d17c3..4b930e16e1994 100644 --- a/src/algebraic_geometry/open_immersion/Scheme.lean +++ b/src/algebraic_geometry/open_immersion/Scheme.lean @@ -10,6 +10,9 @@ import category_theory.limits.shapes.comm_sq /-! # Open immersions of schemes +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + -/ noncomputable theory diff --git a/src/analysis/calculus/implicit.lean b/src/analysis/calculus/implicit.lean index 9e4b3ecc5dca0..21bbda6901e3a 100644 --- a/src/analysis/calculus/implicit.lean +++ b/src/analysis/calculus/implicit.lean @@ -9,6 +9,9 @@ import analysis.normed_space.complemented /-! # Implicit function theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We prove three versions of the implicit function theorem. First we define a structure `implicit_function_data` that holds arguments for the most general version of the implicit function theorem, see `implicit_function_data.implicit_function` diff --git a/src/analysis/matrix.lean b/src/analysis/matrix.lean index aef241badd380..45b429a99d94e 100644 --- a/src/analysis/matrix.lean +++ b/src/analysis/matrix.lean @@ -10,6 +10,9 @@ import analysis.inner_product_space.pi_L2 /-! # Matrices as a normed space +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we provide the following non-instances for norms on matrices: * The elementwise norm: diff --git a/src/analysis/normed_space/matrix_exponential.lean b/src/analysis/normed_space/matrix_exponential.lean index 529ec5ab24174..678d4710a310b 100644 --- a/src/analysis/normed_space/matrix_exponential.lean +++ b/src/analysis/normed_space/matrix_exponential.lean @@ -14,6 +14,9 @@ import topology.uniform_space.matrix /-! # Lemmas about the matrix exponential +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we provide results about `exp` on `matrix`s over a topological or normed algebra. Note that generic results over all topological spaces such as `exp_zero` can be used on matrices without issue, so are not repeated here. The topological results specific to matrices are: diff --git a/src/analysis/normed_space/star/matrix.lean b/src/analysis/normed_space/star/matrix.lean index 035eb8318992f..0de74bb3ddbeb 100644 --- a/src/analysis/normed_space/star/matrix.lean +++ b/src/analysis/normed_space/star/matrix.lean @@ -11,6 +11,9 @@ import linear_algebra.unitary_group /-! # Unitary matrices +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file collects facts about the unitary matrices over `𝕜` (either `ℝ` or `ℂ`). -/ diff --git a/src/data/matrix/invertible.lean b/src/data/matrix/invertible.lean index 0ff393baefcf3..b7ab93f0e6d2f 100644 --- a/src/data/matrix/invertible.lean +++ b/src/data/matrix/invertible.lean @@ -8,6 +8,9 @@ import data.matrix.basic /-! # Extra lemmas about invertible matrices +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Many of the `invertible` lemmas are about `*`; this restates them to be about `⬝`. For lemmas about the matrix inverse in terms of the determinant and adjugate, see `matrix.has_inv` diff --git a/src/linear_algebra/clifford_algebra/basic.lean b/src/linear_algebra/clifford_algebra/basic.lean index f6828941f3658..45eada9e874ed 100644 --- a/src/linear_algebra/clifford_algebra/basic.lean +++ b/src/linear_algebra/clifford_algebra/basic.lean @@ -11,6 +11,9 @@ import linear_algebra.quadratic_form.isometry /-! # Clifford Algebras +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We construct the Clifford algebra of a module `M` over a commutative ring `R`, equipped with a quadratic_form `Q`. diff --git a/src/linear_algebra/exterior_algebra/basic.lean b/src/linear_algebra/exterior_algebra/basic.lean index a5bc276ee5af1..ddde814780643 100644 --- a/src/linear_algebra/exterior_algebra/basic.lean +++ b/src/linear_algebra/exterior_algebra/basic.lean @@ -10,6 +10,9 @@ import linear_algebra.alternating /-! # Exterior Algebras +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We construct the exterior algebra of a module `M` over a commutative semiring `R`. ## Notation diff --git a/src/number_theory/cyclotomic/basic.lean b/src/number_theory/cyclotomic/basic.lean index abb7208658c20..cc16b9b61dbfc 100644 --- a/src/number_theory/cyclotomic/basic.lean +++ b/src/number_theory/cyclotomic/basic.lean @@ -10,6 +10,9 @@ import field_theory.galois /-! # Cyclotomic extensions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Let `A` and `B` be commutative rings with `algebra A B`. For `S : set ℕ+`, we define a class `is_cyclotomic_extension S A B` expressing the fact that `B` is obtained from `A` by adding `n`-th primitive roots of unity, for all `n ∈ S`. diff --git a/src/number_theory/cyclotomic/primitive_roots.lean b/src/number_theory/cyclotomic/primitive_roots.lean index 971334e1bba6b..af57304165afc 100644 --- a/src/number_theory/cyclotomic/primitive_roots.lean +++ b/src/number_theory/cyclotomic/primitive_roots.lean @@ -14,6 +14,9 @@ import ring_theory.polynomial.cyclotomic.expand /-! # Primitive roots in cyclotomic fields + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. If `is_cyclotomic_extension {n} A B`, we define an element `zeta n A B : B` that is a primitive `n`th-root of unity in `B` and we study its properties. We also prove related theorems under the more general assumption of just being a primitive root, for reasons described in the implementation diff --git a/src/number_theory/number_field/embeddings.lean b/src/number_theory/number_field/embeddings.lean index 2e1883a6753cd..a42a2fc245ed5 100644 --- a/src/number_theory/number_field/embeddings.lean +++ b/src/number_theory/number_field/embeddings.lean @@ -13,6 +13,9 @@ import topology.instances.complex /-! # Embeddings of number fields + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. This file defines the embeddings of a number field into an algebraic closed field. ## Main Results diff --git a/src/number_theory/number_field/units.lean b/src/number_theory/number_field/units.lean index 106abb57e6437..26fae65f4afee 100644 --- a/src/number_theory/number_field/units.lean +++ b/src/number_theory/number_field/units.lean @@ -7,6 +7,9 @@ import number_theory.number_field.norm /-! # Units of a number field + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. We prove results about the group `(𝓞 K)ˣ` of units of the ring of integers `𝓞 K` of a number field `K`. diff --git a/src/number_theory/zsqrtd/gaussian_int.lean b/src/number_theory/zsqrtd/gaussian_int.lean index 4c121a2f9f4d5..dc145904b9aaa 100644 --- a/src/number_theory/zsqrtd/gaussian_int.lean +++ b/src/number_theory/zsqrtd/gaussian_int.lean @@ -11,6 +11,9 @@ import ring_theory.principal_ideal_domain /-! # Gaussian integers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The Gaussian integers are complex integer, complex numbers whose real and imaginary parts are both integers. diff --git a/src/ring_theory/dedekind_domain/S_integer.lean b/src/ring_theory/dedekind_domain/S_integer.lean index ec1944af6e824..a3cdfe38b2547 100644 --- a/src/ring_theory/dedekind_domain/S_integer.lean +++ b/src/ring_theory/dedekind_domain/S_integer.lean @@ -9,6 +9,9 @@ import ring_theory.dedekind_domain.adic_valuation /-! # `S`-integers and `S`-units of fraction fields of Dedekind domains +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Let `K` be the field of fractions of a Dedekind domain `R`, and let `S` be a set of prime ideals in the height one spectrum of `R`. An `S`-integer of `K` is defined to have `v`-adic valuation at most one for all primes ideals `v` away from `S`, whereas an `S`-unit of `Kˣ` is defined to have `v`-adic diff --git a/src/ring_theory/dedekind_domain/adic_valuation.lean b/src/ring_theory/dedekind_domain/adic_valuation.lean index 86e0bac6f8b44..a5d737de91673 100644 --- a/src/ring_theory/dedekind_domain/adic_valuation.lean +++ b/src/ring_theory/dedekind_domain/adic_valuation.lean @@ -11,6 +11,9 @@ import algebra.order.group.type_tags /-! # Adic valuations on Dedekind domains + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. Given a Dedekind domain `R` of Krull dimension 1 and a maximal ideal `v` of `R`, we define the `v`-adic valuation on `R` and its extension to the field of fractions `K` of `R`. We prove several properties of this valuation, including the existence of uniformizers. diff --git a/src/ring_theory/dedekind_domain/finite_adele_ring.lean b/src/ring_theory/dedekind_domain/finite_adele_ring.lean index 8c599af57d482..78cf47915dd32 100644 --- a/src/ring_theory/dedekind_domain/finite_adele_ring.lean +++ b/src/ring_theory/dedekind_domain/finite_adele_ring.lean @@ -8,6 +8,9 @@ import ring_theory.dedekind_domain.adic_valuation /-! # The finite adèle ring of a Dedekind domain + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. We define the ring of finite adèles of a Dedekind domain `R`. ## Main definitions diff --git a/src/ring_theory/discriminant.lean b/src/ring_theory/discriminant.lean index 21a559720dab7..57849e8c1f982 100644 --- a/src/ring_theory/discriminant.lean +++ b/src/ring_theory/discriminant.lean @@ -11,6 +11,9 @@ import number_theory.number_field.basic /-! # Discriminant of a family of vectors +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given an `A`-algebra `B` and `b`, an `ι`-indexed family of elements of `B`, we define the *discriminant* of `b` as the determinant of the matrix whose `(i j)`-th element is the trace of `b i * b j`. diff --git a/src/ring_theory/ideal/norm.lean b/src/ring_theory/ideal/norm.lean index b0e0cdd6fcd75..ef4f65f73d77b 100644 --- a/src/ring_theory/ideal/norm.lean +++ b/src/ring_theory/ideal/norm.lean @@ -18,6 +18,9 @@ import ring_theory.localization.norm # Ideal norms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the absolute ideal norm `ideal.abs_norm (I : ideal R) : ℕ` as the cardinality of the quotient `R ⧸ I` (setting it to 0 if the cardinality is infinite), and the relative ideal norm `ideal.span_norm R (I : ideal S) : ideal S` as the ideal spanned by From 93f880918cb51905fd51b76add8273cbc27718ab Mon Sep 17 00:00:00 2001 From: Hanting Zhang Date: Fri, 23 Jun 2023 19:10:46 +0000 Subject: [PATCH 1219/1291] feat(topology/metric_space/dilation): Dilations on metric spaces (#14315) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We define `dilation α β` as the type of maps that satisfy `edist (f x) (f y) = r * edist x y` for all `x y`. Here `r : ℝ≥0`, so we do not exclude the degenerate case of dilations which collapse into constant maps. After this I will extend to `{linear, affine}_dilation_{equiv}`s and `{linear, affine}_isometry_{equiv}`s. Co-authored-by: Yury G. Kudryashov Co-authored-by: JovanGerb --- docs/references.bib | 12 + src/topology/metric_space/dilation.lean | 393 ++++++++++++++++++++++++ 2 files changed, 405 insertions(+) create mode 100644 src/topology/metric_space/dilation.lean diff --git a/docs/references.bib b/docs/references.bib index 2569176a1a9f9..d97389ddd156d 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -159,6 +159,18 @@ @Book{ behrends1979 zbl = {0436.46013} } +@Book{ berger1987, + author = {Marcel Berger}, + title = {Geometry I}, + publisher = {Springer Berlin}, + year = 1987, + issn = {0172-5939}, + pages = {XIV, 432}, + series = {Universitext}, + address = {Heidelberg}, + edition = 1 +} + @Article{ bernstein1912, author = {Bernstein, S.}, year = {1912}, diff --git a/src/topology/metric_space/dilation.lean b/src/topology/metric_space/dilation.lean new file mode 100644 index 0000000000000..24195f12cdfcb --- /dev/null +++ b/src/topology/metric_space/dilation.lean @@ -0,0 +1,393 @@ +/- +Copyright (c) 2022 Hanting Zhang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Dilations of emetric and metric spaces +Authors: Hanting Zhang +-/ +import topology.metric_space.antilipschitz +import data.fun_like.basic + +/-! +# Dilations + +We define dilations, i.e., maps between emetric spaces that satisfy +`edist (f x) (f y) = r * edist x y` for some `r ∉ {0, ∞}`. + +The value `r = 0` is not allowed because we want dilations of (e)metric spaces to be automatically +injective. The value `r = ∞` is not allowed because this way we can define `dilation.ratio f : ℝ≥0`, +not `dilation.ratio f : ℝ≥0∞`. Also, we do not often need maps sending distinct points to points at +infinite distance. + +## Main defintions + +* `dilation.ratio f : ℝ≥0`: the value of `r` in the relation above, defaulting to 1 in the case + where it is not well-defined. + +## Implementation notes + +The type of dilations defined in this file are also referred to as "similarities" or "similitudes" +by other authors. The name `dilation` was choosen to match the Wikipedia name. + +Since a lot of elementary properties don't require `eq_of_dist_eq_zero` we start setting up the +theory for `pseudo_emetric_space` and we specialize to `pseudo_metric_space` and `metric_space` when +needed. + +## TODO + +- Introduce dilation equivs. +- Refactor the `isometry` API to match the `*_hom_class` API below. + +## References + +- https://en.wikipedia.org/wiki/Dilation_(metric_space) +- [Marcel Berger, *Geometry*][berger1987] +-/ + +noncomputable theory + +open function set +open_locale topology ennreal nnreal classical + +section defs + +variables (α : Type*) (β : Type*) [pseudo_emetric_space α] [pseudo_emetric_space β] + +/-- A dilation is a map that uniformly scales the edistance between any two points. -/ +structure dilation := +(to_fun : α → β) +(edist_eq' : ∃ r : ℝ≥0, r ≠ 0 ∧ ∀ x y : α, edist (to_fun x) (to_fun y) = r * edist x y) + +/-- +`dilation_class F α β r` states that `F` is a type of `r`-dilations. +You should extend this typeclass when you extend `dilation`. +-/ +class dilation_class (F : Type*) (α β : out_param $ Type*) + [pseudo_emetric_space α] [pseudo_emetric_space β] extends fun_like F α (λ _, β) := +(edist_eq' : ∀ (f : F), ∃ r : ℝ≥0, r ≠ 0 ∧ ∀ (x y : α), edist (f x) (f y) = r * edist x y) + +end defs + +namespace dilation +variables {α : Type*} {β : Type*} {γ : Type*} {F : Type*} {G : Type*} + +section setup +variables [pseudo_emetric_space α] [pseudo_emetric_space β] + +instance to_dilation_class : + dilation_class (dilation α β) α β := +{ coe := to_fun, + coe_injective' := λ f g h, by { cases f; cases g; congr', }, + edist_eq' := λ f, edist_eq' f } + +instance : has_coe_to_fun (dilation α β) (λ _, α → β) := fun_like.has_coe_to_fun + +@[simp] lemma to_fun_eq_coe {f : dilation α β} : f.to_fun = (f : α → β) := rfl + +@[simp] lemma coe_mk (f : α → β) (h) : ⇑(⟨f, h⟩ : dilation α β) = f := rfl + +lemma congr_fun {f g : dilation α β} (h : f = g) (x : α) : f x = g x := fun_like.congr_fun h x +lemma congr_arg (f : dilation α β) {x y : α} (h : x = y) : f x = f y := fun_like.congr_arg f h + +@[ext] theorem ext {f g : dilation α β} (h : ∀ x, f x = g x) : f = g := +fun_like.ext f g h + +lemma ext_iff {f g : dilation α β} : f = g ↔ ∀ x, f x = g x := fun_like.ext_iff + +@[simp] lemma mk_coe (f : dilation α β) (h) : dilation.mk f h = f := ext $ λ _, rfl + +/-- Copy of a `dilation` with a new `to_fun` equal to the old one. Useful to fix definitional +equalities. -/ +@[simps { fully_applied := ff }] +protected def copy (f : dilation α β) (f' : α → β) (h : f' = ⇑f) : dilation α β := +{ to_fun := f', + edist_eq' := h.symm ▸ f.edist_eq' } + +lemma copy_eq_self (f : dilation α β) {f' : α → β} (h : f' = f) : f.copy f' h = f := +fun_like.ext' h + +/-- The ratio of a dilation `f`. If the ratio is undefined (i.e., the distance between any two +points in `α` is either zero or infinity), then we choose one as the ratio. -/ +def ratio [dilation_class F α β] (f : F) : ℝ≥0 := +if ∀ x y : α, edist x y = 0 ∨ edist x y = ⊤ then 1 else (dilation_class.edist_eq' f).some + +lemma ratio_ne_zero [dilation_class F α β] (f : F) : ratio f ≠ 0 := +begin + rw ratio, split_ifs, + { exact one_ne_zero, }, + exact (dilation_class.edist_eq' f).some_spec.1, +end + +lemma ratio_pos [dilation_class F α β] (f : F) : 0 < ratio f := +(ratio_ne_zero f).bot_lt + +@[simp] lemma edist_eq [dilation_class F α β] (f : F) (x y : α) : + edist (f x) (f y) = ratio f * edist x y := +begin + rw ratio, split_ifs with key, + { rcases dilation_class.edist_eq' f with ⟨r, hne, hr⟩, + replace hr := hr x y, + cases key x y, + { simp only [hr, h, mul_zero] }, + { simp [hr, h, hne] } }, + exact (dilation_class.edist_eq' f).some_spec.2 x y, +end + +@[simp] lemma nndist_eq {α β F : Type*} [pseudo_metric_space α] [pseudo_metric_space β] + [dilation_class F α β] (f : F) (x y : α) : nndist (f x) (f y) = ratio f * nndist x y := +by simp only [← ennreal.coe_eq_coe, ← edist_nndist, ennreal.coe_mul, edist_eq] + +@[simp] lemma dist_eq {α β F : Type*} [pseudo_metric_space α] [pseudo_metric_space β] + [dilation_class F α β] (f : F) (x y : α) : dist (f x) (f y) = ratio f * dist x y := +by simp only [dist_nndist, nndist_eq, nnreal.coe_mul] + +/-- The `ratio` is equal to the distance ratio for any two points with nonzero finite distance. +`dist` and `nndist` versions below -/ +lemma ratio_unique [dilation_class F α β] {f : F} {x y : α} {r : ℝ≥0} + (h₀ : edist x y ≠ 0) (htop : edist x y ≠ ⊤) (hr : edist (f x) (f y) = r * edist x y) : + r = ratio f := +by simpa only [hr, ennreal.mul_eq_mul_right h₀ htop, ennreal.coe_eq_coe] using edist_eq f x y + +/-- The `ratio` is equal to the distance ratio for any two points +with nonzero finite distance; `nndist` version -/ +lemma ratio_unique_of_nndist_ne_zero {α β F : Type*} [pseudo_metric_space α] [pseudo_metric_space β] + [dilation_class F α β] {f : F} {x y : α} {r : ℝ≥0} (hxy : nndist x y ≠ 0) + (hr : nndist (f x) (f y) = r * nndist x y) : r = ratio f := +ratio_unique (by rwa [edist_nndist, ennreal.coe_ne_zero]) (edist_ne_top x y) + (by rw [edist_nndist, edist_nndist, hr, ennreal.coe_mul]) + +/-- The `ratio` is equal to the distance ratio for any two points +with nonzero finite distance; `dist` version -/ +lemma ratio_unique_of_dist_ne_zero {α β} {F : Type*} [pseudo_metric_space α] [pseudo_metric_space β] + [dilation_class F α β] {f : F} {x y : α} {r : ℝ≥0} + (hxy : dist x y ≠ 0) (hr : dist (f x) (f y) = r * dist x y) : + r = ratio f := +ratio_unique_of_nndist_ne_zero (nnreal.coe_ne_zero.1 hxy) $ nnreal.eq $ + by rw [coe_nndist, hr, nnreal.coe_mul, coe_nndist] + +/-- Alternative `dilation` constructor when the distance hypothesis is over `nndist` -/ +def mk_of_nndist_eq {α β} + [pseudo_metric_space α] [pseudo_metric_space β] + (f : α → β) (h : ∃ (r : ℝ≥0), r ≠ 0 ∧ ∀ (x y : α), nndist (f x) (f y) = r * nndist x y) : + dilation α β := +{ to_fun := f, + edist_eq' := + begin + rcases h with ⟨r, hne, h⟩, + refine ⟨r, hne, λ x y, _⟩, + rw [edist_nndist, edist_nndist, ← ennreal.coe_mul, h x y], + end } + +@[simp] lemma coe_mk_of_nndist_eq {α β} + [pseudo_metric_space α] [pseudo_metric_space β] + (f : α → β) (h) : ⇑(mk_of_nndist_eq f h : dilation α β) = f := rfl + +@[simp] lemma mk_coe_of_nndist_eq {α β} + [pseudo_metric_space α] [pseudo_metric_space β] + (f : dilation α β) (h) : dilation.mk_of_nndist_eq f h = f := +ext $ λ _, rfl + +/-- Alternative `dilation` constructor when the distance hypothesis is over `dist` -/ +def mk_of_dist_eq {α β} + [pseudo_metric_space α] [pseudo_metric_space β] + (f : α → β) (h : ∃ (r : ℝ≥0), r ≠ 0 ∧ ∀ (x y : α), dist (f x) (f y) = r * dist x y) : + dilation α β := +mk_of_nndist_eq f $ h.imp $ λ r hr, + ⟨hr.1, λ x y, nnreal.eq $ by rw [coe_nndist, hr.2, nnreal.coe_mul, coe_nndist]⟩ + +@[simp] lemma coe_mk_of_dist_eq {α β} + [pseudo_metric_space α] [pseudo_metric_space β] + (f : α → β) (h) : ⇑(mk_of_dist_eq f h : dilation α β) = f := rfl + +@[simp] lemma mk_coe_of_dist_eq {α β} + [pseudo_metric_space α] [pseudo_metric_space β] + (f : dilation α β) (h) : dilation.mk_of_dist_eq f h = f := +ext $ λ _, rfl + +end setup + +section pseudo_emetric_dilation +variables [pseudo_emetric_space α] [pseudo_emetric_space β] [pseudo_emetric_space γ] +variables [dilation_class F α β] [dilation_class G β γ] +variables (f : F) (g : G) {x y z : α} {s : set α} + +lemma lipschitz : lipschitz_with (ratio f) (f : α → β) := λ x y, (edist_eq f x y).le + +lemma antilipschitz : antilipschitz_with (ratio f)⁻¹ (f : α → β) := +λ x y, have hr : ratio f ≠ 0 := ratio_ne_zero f, by exact_mod_cast + (ennreal.mul_le_iff_le_inv (ennreal.coe_ne_zero.2 hr) ennreal.coe_ne_top).1 (edist_eq f x y).ge + +/-- A dilation from an emetric space is injective -/ +protected lemma injective {α : Type*} [emetric_space α] [dilation_class F α β] (f : F) : + injective f := (antilipschitz f).injective + +/-- The identity is a dilation -/ +protected def id (α) [pseudo_emetric_space α] : dilation α α := +{ to_fun := _root_.id, + edist_eq' := ⟨1, one_ne_zero, λ x y, by simp only [id.def, ennreal.coe_one, one_mul]⟩ } + +instance : inhabited (dilation α α) := ⟨dilation.id α⟩ + +@[simp, protected] lemma coe_id : ⇑(dilation.id α) = id := rfl + +lemma id_ratio : ratio (dilation.id α) = 1 := +begin + by_cases h : ∀ x y : α, edist x y = 0 ∨ edist x y = ∞, + { rw [ratio, if_pos h] }, + { push_neg at h, + rcases h with ⟨x, y, hne⟩, + refine (ratio_unique hne.1 hne.2 _).symm, + simp } +end + +/-- The composition of dilations is a dilation -/ +def comp (g : dilation β γ) (f : dilation α β) : dilation α γ := +{ to_fun := g ∘ f, + edist_eq' := ⟨ratio g * ratio f, + mul_ne_zero (ratio_ne_zero g) (ratio_ne_zero f), + λ x y, by { simp only [edist_eq, ennreal.coe_mul], ring, }⟩ } + +lemma comp_assoc {δ : Type*} [pseudo_emetric_space δ] + (f : dilation α β) (g : dilation β γ) (h : dilation γ δ) : + (h.comp g).comp f = h.comp (g.comp f) := rfl + +@[simp] lemma coe_comp (g : dilation β γ) (f : dilation α β) : + (g.comp f : α → γ) = g ∘ f := rfl + +lemma comp_apply (g : dilation β γ) (f : dilation α β) (x : α) : + (g.comp f : α → γ) x = (g (f x)) := rfl + +/-- Ratio of the composition `g.comp f` of two dilations is the product of their ratios. We assume +that the domain `α` of `f` is nontrivial, otherwise `ratio f = ratio (g.comp f) = 1` but `ratio g` +may have any value. -/ +@[simp] lemma comp_ratio + {g : dilation β γ} {f : dilation α β} (hne : ∃ x y : α, edist x y ≠ 0 ∧ edist x y ≠ ⊤) : + ratio (g.comp f) = ratio g * ratio f := +begin + rcases hne with ⟨x, y, hα⟩, + have hgf := (edist_eq (g.comp f) x y).symm, + simp only [dist_eq, coe_comp, ← mul_assoc, mul_eq_mul_right_iff] at hgf, + rw [edist_eq, edist_eq, ← mul_assoc, ennreal.mul_eq_mul_right hα.1 hα.2] at hgf, + rwa [← ennreal.coe_eq_coe, ennreal.coe_mul], +end + +@[simp] lemma comp_id (f : dilation α β) : f.comp (dilation.id α) = f := ext $ λ x, rfl + +@[simp] lemma id_comp (f : dilation α β) : (dilation.id β).comp f = f := ext $ λ x, rfl + +instance : monoid (dilation α α) := +{ one := dilation.id α, + mul := comp, + mul_one := comp_id, + one_mul := id_comp, + mul_assoc := λ f g h, comp_assoc _ _ _ } + +lemma one_def : (1 : dilation α α) = dilation.id α := rfl +lemma mul_def (f g : dilation α α) : f * g = f.comp g := rfl + +@[simp] lemma coe_one : ⇑(1 : dilation α α) = _root_.id := rfl +@[simp] lemma coe_mul (f g : dilation α α) : ⇑(f * g) = f ∘ g := rfl + +lemma cancel_right {g₁ g₂ : dilation β γ} {f : dilation α β} (hf : surjective f) : + g₁.comp f = g₂.comp f ↔ g₁ = g₂ := +⟨λ h, dilation.ext $ hf.forall.2 (ext_iff.1 h), λ h, h ▸ rfl⟩ + +lemma cancel_left {g : dilation β γ} {f₁ f₂ : dilation α β} (hg : injective g) : + g.comp f₁ = g.comp f₂ ↔ f₁ = f₂ := +⟨λ h, dilation.ext $ λ x, hg $ by rw [← comp_apply, h, comp_apply], λ h, h ▸ rfl⟩ + +/-- A dilation from a metric space is a uniform inducing map -/ +protected theorem uniform_inducing : uniform_inducing (f : α → β) := +(antilipschitz f).uniform_inducing (lipschitz f).uniform_continuous + +lemma tendsto_nhds_iff {ι : Type*} {g : ι → α} {a : filter ι} {b : α} : + filter.tendsto g a (𝓝 b) ↔ filter.tendsto ((f : α → β) ∘ g) a (𝓝 (f b)) := +(dilation.uniform_inducing f).inducing.tendsto_nhds_iff + +/-- A dilation is continuous. -/ +lemma to_continuous : continuous (f : α → β) := +(lipschitz f).continuous + +/-- Dilations scale the diameter by `ratio f` in pseudoemetric spaces. -/ +lemma ediam_image (s : set α) : + emetric.diam ((f : α → β) '' s) = ratio f * emetric.diam s := +begin + refine ((lipschitz f).ediam_image_le s).antisymm _, + apply ennreal.mul_le_of_le_div', + rw [div_eq_mul_inv, mul_comm, ← ennreal.coe_inv], + exacts [(antilipschitz f).le_mul_ediam_image s, ratio_ne_zero f], +end + +/-- A dilation scales the diameter of the range by `ratio f`. -/ +lemma ediam_range : + emetric.diam (range (f : α → β)) = ratio f * emetric.diam (univ : set α) := +by { rw ← image_univ, exact ediam_image f univ } + +/-- A dilation maps balls to balls and scales the radius by `ratio f`. -/ +lemma maps_to_emetric_ball (x : α) (r : ℝ≥0∞) : + maps_to (f : α → β) (emetric.ball x r) (emetric.ball (f x) (ratio f * r)) := +λ y hy, (edist_eq f y x).trans_lt $ + (ennreal.mul_lt_mul_left (ennreal.coe_ne_zero.2 $ ratio_ne_zero f) ennreal.coe_ne_top).2 hy + +/-- A dilation maps closed balls to closed balls and scales the radius by `ratio f`. -/ +lemma maps_to_emetric_closed_ball (x : α) (r' : ℝ≥0∞) : + maps_to (f : α → β) (emetric.closed_ball x r') (emetric.closed_ball (f x) (ratio f * r')) := +λ y hy, (edist_eq f y x).trans_le $ mul_le_mul_left' hy _ + +lemma comp_continuous_on_iff {γ} [topological_space γ] {g : γ → α} {s : set γ} : + continuous_on ((f : α → β) ∘ g) s ↔ continuous_on g s := +(dilation.uniform_inducing f).inducing.continuous_on_iff.symm + +lemma comp_continuous_iff {γ} [topological_space γ] {g : γ → α} : + continuous ((f : α → β) ∘ g) ↔ continuous g := +(dilation.uniform_inducing f).inducing.continuous_iff.symm + +end pseudo_emetric_dilation --section + +section emetric_dilation +variables [emetric_space α] + +/-- A dilation from a metric space is a uniform embedding -/ +protected theorem uniform_embedding [pseudo_emetric_space β] [dilation_class F α β] + (f : F) : uniform_embedding f := +(antilipschitz f).uniform_embedding (lipschitz f).uniform_continuous + +/-- A dilation from a metric space is an embedding -/ +protected theorem embedding [pseudo_emetric_space β] [dilation_class F α β] + (f : F) : embedding (f : α → β) := +(dilation.uniform_embedding f).embedding + +/-- A dilation from a complete emetric space is a closed embedding -/ +protected theorem closed_embedding [complete_space α] [emetric_space β] [dilation_class F α β] + (f : F) : closed_embedding f := +(antilipschitz f).closed_embedding (lipschitz f).uniform_continuous + +end emetric_dilation --section + +section pseudo_metric_dilation +variables [pseudo_metric_space α] [pseudo_metric_space β] [dilation_class F α β] (f : F) + +/-- A dilation scales the diameter by `ratio f` in pseudometric spaces. -/ +lemma diam_image (s : set α) : metric.diam ((f : α → β) '' s) = ratio f * metric.diam s := +by { simp [metric.diam, ediam_image, ennreal.to_real_mul], } + +lemma diam_range : metric.diam (range (f : α → β)) = ratio f * metric.diam (univ : set α) := +by rw [← image_univ, diam_image] + +/-- A dilation maps balls to balls and scales the radius by `ratio f`. -/ +lemma maps_to_ball (x : α) (r' : ℝ) : + maps_to (f : α → β) (metric.ball x r') (metric.ball (f x) (ratio f * r')) := +λ y hy, (dist_eq f y x).trans_lt $ (mul_lt_mul_left $ nnreal.coe_pos.2 $ ratio_pos f).2 hy + +/-- A dilation maps spheres to spheres and scales the radius by `ratio f`. -/ +lemma maps_to_sphere (x : α) (r' : ℝ) : + maps_to (f : α → β) (metric.sphere x r') (metric.sphere (f x) (ratio f * r')) := +λ y hy, metric.mem_sphere.mp hy ▸ dist_eq f y x + +/-- A dilation maps closed balls to closed balls and scales the radius by `ratio f`. -/ +lemma maps_to_closed_ball (x : α) (r' : ℝ) : + maps_to (f : α → β) (metric.closed_ball x r') (metric.closed_ball (f x) (ratio f * r')) := +λ y hy, (dist_eq f y x).trans_le $ mul_le_mul_of_nonneg_left hy (nnreal.coe_nonneg _) + +end pseudo_metric_dilation -- section + +end dilation From fdc286cc6967a012f41b87f76dcd2797b53152af Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Sat, 24 Jun 2023 06:34:52 +0000 Subject: [PATCH 1220/1291] chore(*): add mathlib4 synchronization comments (#19214) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebraic_geometry.presheafed_space.gluing` * `analysis.calculus.bump_function_findim` * `field_theory.polynomial_galois_group` * `geometry.manifold.sheaf.basic` * `linear_algebra.clifford_algebra.conjugation` * `linear_algebra.clifford_algebra.grading` * `linear_algebra.clifford_algebra.star` * `linear_algebra.free_module.norm` * `number_theory.cyclotomic.discriminant` * `number_theory.legendre_symbol.add_character` * `number_theory.number_field.canonical_embedding` --- src/algebraic_geometry/presheafed_space/gluing.lean | 3 +++ src/analysis/calculus/bump_function_findim.lean | 3 +++ src/field_theory/polynomial_galois_group.lean | 3 +++ src/geometry/manifold/sheaf/basic.lean | 3 +++ src/linear_algebra/clifford_algebra/conjugation.lean | 3 +++ src/linear_algebra/clifford_algebra/grading.lean | 3 +++ src/linear_algebra/clifford_algebra/star.lean | 3 +++ src/linear_algebra/free_module/norm.lean | 3 +++ src/number_theory/cyclotomic/discriminant.lean | 3 +++ src/number_theory/legendre_symbol/add_character.lean | 3 +++ src/number_theory/number_field/canonical_embedding.lean | 3 +++ 11 files changed, 33 insertions(+) diff --git a/src/algebraic_geometry/presheafed_space/gluing.lean b/src/algebraic_geometry/presheafed_space/gluing.lean index 8f214a1122052..c964bb9f338bf 100644 --- a/src/algebraic_geometry/presheafed_space/gluing.lean +++ b/src/algebraic_geometry/presheafed_space/gluing.lean @@ -10,6 +10,9 @@ import algebraic_geometry.locally_ringed_space.has_colimits /-! # Gluing Structured spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given a family of gluing data of structured spaces (presheafed spaces, sheafed spaces, or locally ringed spaces), we may glue them together. diff --git a/src/analysis/calculus/bump_function_findim.lean b/src/analysis/calculus/bump_function_findim.lean index 67a52967d0b5a..b46cf1c50c964 100644 --- a/src/analysis/calculus/bump_function_findim.lean +++ b/src/analysis/calculus/bump_function_findim.lean @@ -12,6 +12,9 @@ import data.set.pointwise.support /-! # Bump functions in finite-dimensional vector spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Let `E` be a finite-dimensional real normed vector space. We show that any open set `s` in `E` is exactly the support of a smooth function taking values in `[0, 1]`, in `is_open.exists_smooth_support_eq`. diff --git a/src/field_theory/polynomial_galois_group.lean b/src/field_theory/polynomial_galois_group.lean index d25a4855e0602..d522d1b6ba71f 100644 --- a/src/field_theory/polynomial_galois_group.lean +++ b/src/field_theory/polynomial_galois_group.lean @@ -10,6 +10,9 @@ import group_theory.perm.cycle.type /-! # Galois Groups of Polynomials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we introduce the Galois group of a polynomial `p` over a field `F`, defined as the automorphism group of its splitting field. We also provide some results about some extension `E` above `p.splitting_field`, and some specific diff --git a/src/geometry/manifold/sheaf/basic.lean b/src/geometry/manifold/sheaf/basic.lean index 35fe4545d0d01..5c4722153c102 100644 --- a/src/geometry/manifold/sheaf/basic.lean +++ b/src/geometry/manifold/sheaf/basic.lean @@ -8,6 +8,9 @@ import topology.sheaves.local_predicate /-! # Generic construction of a sheaf from a `local_invariant_prop` on a manifold +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file constructs the sheaf-of-types of functions `f : M → M'` (for charted spaces `M`, `M'`) which satisfy the lifted property `lift_prop P` associated to some locally invariant (in the sense of `structure_groupoid.local_invariant_prop`) property `P` on the model spaces of `M` and `M'`. For diff --git a/src/linear_algebra/clifford_algebra/conjugation.lean b/src/linear_algebra/clifford_algebra/conjugation.lean index 07f8f5174fd87..f25611184e923 100644 --- a/src/linear_algebra/clifford_algebra/conjugation.lean +++ b/src/linear_algebra/clifford_algebra/conjugation.lean @@ -9,6 +9,9 @@ import algebra.module.opposites /-! # Conjugations +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the grade reversal and grade involution functions on multivectors, `reverse` and `involute`. Together, these operations compose to form the "Clifford conjugate", hence the name of this file. diff --git a/src/linear_algebra/clifford_algebra/grading.lean b/src/linear_algebra/clifford_algebra/grading.lean index d32e266c568e9..6e9f3a429a818 100644 --- a/src/linear_algebra/clifford_algebra/grading.lean +++ b/src/linear_algebra/clifford_algebra/grading.lean @@ -10,6 +10,9 @@ import ring_theory.graded_algebra.basic /-! # Results about the grading structure of the clifford algebra +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The main result is `clifford_algebra.graded_algebra`, which says that the clifford algebra is a ℤ₂-graded algebra (or "superalgebra"). -/ diff --git a/src/linear_algebra/clifford_algebra/star.lean b/src/linear_algebra/clifford_algebra/star.lean index 9c340fcbcb04f..697a4b2f4464d 100644 --- a/src/linear_algebra/clifford_algebra/star.lean +++ b/src/linear_algebra/clifford_algebra/star.lean @@ -8,6 +8,9 @@ import linear_algebra.clifford_algebra.conjugation /-! # Star structure on `clifford_algebra` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the "clifford conjugation", equal to `reverse (involute x)`, and assigns it the `star` notation. diff --git a/src/linear_algebra/free_module/norm.lean b/src/linear_algebra/free_module/norm.lean index 7f73836649fac..f9366158ef8a6 100644 --- a/src/linear_algebra/free_module/norm.lean +++ b/src/linear_algebra/free_module/norm.lean @@ -9,6 +9,9 @@ import ring_theory.norm /-! # Norms on free modules over principal ideal domains + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ open ideal polynomial diff --git a/src/number_theory/cyclotomic/discriminant.lean b/src/number_theory/cyclotomic/discriminant.lean index f92f560b96935..1875990e97556 100644 --- a/src/number_theory/cyclotomic/discriminant.lean +++ b/src/number_theory/cyclotomic/discriminant.lean @@ -9,6 +9,9 @@ import ring_theory.discriminant /-! # Discriminant of cyclotomic fields + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. We compute the discriminant of a `p ^ n`-th cyclotomic extension. ## Main results diff --git a/src/number_theory/legendre_symbol/add_character.lean b/src/number_theory/legendre_symbol/add_character.lean index e35db8da04b63..47de7d6d7729b 100644 --- a/src/number_theory/legendre_symbol/add_character.lean +++ b/src/number_theory/legendre_symbol/add_character.lean @@ -9,6 +9,9 @@ import field_theory.finite.trace /-! # Additive characters of finite rings and fields +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Let `R` be a finite commutative ring. An *additive character* of `R` with values in another commutative ring `R'` is simply a morphism from the additive group of `R` into the multiplicative monoid of `R'`. diff --git a/src/number_theory/number_field/canonical_embedding.lean b/src/number_theory/number_field/canonical_embedding.lean index bdcc13bebd7ef..d3b0066683505 100644 --- a/src/number_theory/number_field/canonical_embedding.lean +++ b/src/number_theory/number_field/canonical_embedding.lean @@ -8,6 +8,9 @@ import number_theory.number_field.embeddings /-! # Canonical embedding of a number field +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The canonical embedding of a number field `K` of signature `(r₁, r₂)` is the ring homomorphism `K →+* ℝ^r₁ × ℂ^r₂` that sends `x ∈ K` to `(φ_₁(x),...,φ_r₁(x)) × (ψ_₁(x),..., ψ_r₂(x))` where `φ_₁,...,φ_r₁` are its real embeddings and `ψ_₁,..., ψ_r₂` are its complex embeddings (up to From 30faa0c3618ce1472bf6305ae0e3fa56affa3f95 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Mon, 26 Jun 2023 04:51:08 +0000 Subject: [PATCH 1221/1291] chore(*): add mathlib4 synchronization comments (#19215) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `field_theory.abel_ruffini` * `geometry.manifold.algebra.lie_group` * `geometry.manifold.algebra.monoid` * `geometry.manifold.algebra.structures` * `geometry.manifold.bump_function` * `geometry.manifold.cont_mdiff` * `geometry.manifold.cont_mdiff_map` * `geometry.manifold.vector_bundle.basic` * `geometry.manifold.vector_bundle.fiberwise_linear` * `geometry.manifold.vector_bundle.pullback` * `geometry.manifold.vector_bundle.tangent` * `linear_algebra.clifford_algebra.equivs` * `linear_algebra.clifford_algebra.fold` * `number_theory.class_number.finite` * `number_theory.class_number.function_field` * `number_theory.cyclotomic.gal` * `number_theory.cyclotomic.rat` * `number_theory.number_field.class_number` * `representation_theory.group_cohomology.resolution` * `topology.metric_space.dilation` --- src/field_theory/abel_ruffini.lean | 3 +++ src/geometry/manifold/algebra/lie_group.lean | 3 +++ src/geometry/manifold/algebra/monoid.lean | 3 +++ src/geometry/manifold/algebra/structures.lean | 3 +++ src/geometry/manifold/bump_function.lean | 3 +++ src/geometry/manifold/cont_mdiff.lean | 3 +++ src/geometry/manifold/cont_mdiff_map.lean | 3 +++ src/geometry/manifold/vector_bundle/basic.lean | 3 +++ src/geometry/manifold/vector_bundle/fiberwise_linear.lean | 3 +++ src/geometry/manifold/vector_bundle/pullback.lean | 3 +++ src/geometry/manifold/vector_bundle/tangent.lean | 3 +++ src/linear_algebra/clifford_algebra/equivs.lean | 3 +++ src/linear_algebra/clifford_algebra/fold.lean | 3 +++ src/number_theory/class_number/finite.lean | 3 +++ src/number_theory/class_number/function_field.lean | 3 +++ src/number_theory/cyclotomic/gal.lean | 3 +++ src/number_theory/cyclotomic/rat.lean | 3 +++ src/number_theory/number_field/class_number.lean | 3 +++ src/representation_theory/group_cohomology/resolution.lean | 3 +++ src/topology/metric_space/dilation.lean | 3 +++ 20 files changed, 60 insertions(+) diff --git a/src/field_theory/abel_ruffini.lean b/src/field_theory/abel_ruffini.lean index aa5760145959a..8c54b0d3de584 100644 --- a/src/field_theory/abel_ruffini.lean +++ b/src/field_theory/abel_ruffini.lean @@ -11,6 +11,9 @@ import ring_theory.roots_of_unity.basic /-! # The Abel-Ruffini Theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves one direction of the Abel-Ruffini theorem, namely that if an element is solvable by radicals, then its minimal polynomial has solvable Galois group. diff --git a/src/geometry/manifold/algebra/lie_group.lean b/src/geometry/manifold/algebra/lie_group.lean index 32933bde56c91..90009628761f9 100644 --- a/src/geometry/manifold/algebra/lie_group.lean +++ b/src/geometry/manifold/algebra/lie_group.lean @@ -9,6 +9,9 @@ import geometry.manifold.algebra.monoid /-! # Lie groups +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A Lie group is a group that is also a smooth manifold, in which the group operations of multiplication and inversion are smooth maps. Smoothness of the group multiplication means that multiplication is a smooth mapping of the product manifold `G` × `G` into `G`. diff --git a/src/geometry/manifold/algebra/monoid.lean b/src/geometry/manifold/algebra/monoid.lean index 79d38983c7054..3c2be5411253f 100644 --- a/src/geometry/manifold/algebra/monoid.lean +++ b/src/geometry/manifold/algebra/monoid.lean @@ -8,6 +8,9 @@ import geometry.manifold.cont_mdiff_map /-! # Smooth monoid + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. A smooth monoid is a monoid that is also a smooth manifold, in which multiplication is a smooth map of the product manifold `G` × `G` into `G`. diff --git a/src/geometry/manifold/algebra/structures.lean b/src/geometry/manifold/algebra/structures.lean index b8fad067971af..e90599f0c066d 100644 --- a/src/geometry/manifold/algebra/structures.lean +++ b/src/geometry/manifold/algebra/structures.lean @@ -8,6 +8,9 @@ import geometry.manifold.algebra.lie_group /-! # Smooth structures +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define smooth structures that build on Lie groups. We prefer using the term smooth instead of Lie mainly because Lie ring has currently another use in mathematics. -/ diff --git a/src/geometry/manifold/bump_function.lean b/src/geometry/manifold/bump_function.lean index 80262412fa133..c301d8103d3f4 100644 --- a/src/geometry/manifold/bump_function.lean +++ b/src/geometry/manifold/bump_function.lean @@ -9,6 +9,9 @@ import geometry.manifold.cont_mdiff /-! # Smooth bump functions on a smooth manifold +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define `smooth_bump_function I c` to be a bundled smooth "bump" function centered at `c`. It is a structure that consists of two real numbers `0 < r < R` with small enough `R`. We define a coercion to function for this type, and for `f : smooth_bump_function I c`, the function diff --git a/src/geometry/manifold/cont_mdiff.lean b/src/geometry/manifold/cont_mdiff.lean index e6205468df232..2ffd0e1c3c7d8 100644 --- a/src/geometry/manifold/cont_mdiff.lean +++ b/src/geometry/manifold/cont_mdiff.lean @@ -9,6 +9,9 @@ import geometry.manifold.local_invariant_properties /-! # Smooth functions between smooth manifolds +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define `Cⁿ` functions between smooth manifolds, as functions which are `Cⁿ` in charts, and prove basic properties of these notions. diff --git a/src/geometry/manifold/cont_mdiff_map.lean b/src/geometry/manifold/cont_mdiff_map.lean index 9b736749c8f52..e5c95ee2813fd 100644 --- a/src/geometry/manifold/cont_mdiff_map.lean +++ b/src/geometry/manifold/cont_mdiff_map.lean @@ -10,6 +10,9 @@ import topology.continuous_function.basic /-! # Smooth bundled map +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define the type `cont_mdiff_map` of `n` times continuously differentiable bundled maps. -/ diff --git a/src/geometry/manifold/vector_bundle/basic.lean b/src/geometry/manifold/vector_bundle/basic.lean index 5bef774d46b7f..49445ccc558dc 100644 --- a/src/geometry/manifold/vector_bundle/basic.lean +++ b/src/geometry/manifold/vector_bundle/basic.lean @@ -8,6 +8,9 @@ import topology.vector_bundle.constructions /-! # Smooth vector bundles +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines smooth vector bundles over a smooth manifold. Let `E` be a topological vector bundle, with model fiber `F` and base space `B`. We consider `E` as diff --git a/src/geometry/manifold/vector_bundle/fiberwise_linear.lean b/src/geometry/manifold/vector_bundle/fiberwise_linear.lean index cd9e5a761d63a..d2253f00b568d 100644 --- a/src/geometry/manifold/vector_bundle/fiberwise_linear.lean +++ b/src/geometry/manifold/vector_bundle/fiberwise_linear.lean @@ -7,6 +7,9 @@ import geometry.manifold.cont_mdiff /-! # The groupoid of smooth, fiberwise-linear maps +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains preliminaries for the definition of a smooth vector bundle: an associated `structure_groupoid`, the groupoid of `smooth_fiberwise_linear` functions. -/ diff --git a/src/geometry/manifold/vector_bundle/pullback.lean b/src/geometry/manifold/vector_bundle/pullback.lean index c65bc33987d01..44c3c92709f7c 100644 --- a/src/geometry/manifold/vector_bundle/pullback.lean +++ b/src/geometry/manifold/vector_bundle/pullback.lean @@ -8,6 +8,9 @@ import geometry.manifold.vector_bundle.basic /-! # Pullbacks of smooth vector bundles +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines pullbacks of smooth vector bundles over a smooth manifold. ## Main definitions diff --git a/src/geometry/manifold/vector_bundle/tangent.lean b/src/geometry/manifold/vector_bundle/tangent.lean index f5cd27d8429a8..55d597435cb96 100644 --- a/src/geometry/manifold/vector_bundle/tangent.lean +++ b/src/geometry/manifold/vector_bundle/tangent.lean @@ -8,6 +8,9 @@ import geometry.manifold.vector_bundle.basic /-! # Tangent bundles +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the tangent bundle as a smooth vector bundle. Let `M` be a smooth manifold with corners with model `I` on `(E, H)`. We define the tangent bundle diff --git a/src/linear_algebra/clifford_algebra/equivs.lean b/src/linear_algebra/clifford_algebra/equivs.lean index fbfc724ea7963..2c5bf524e78a8 100644 --- a/src/linear_algebra/clifford_algebra/equivs.lean +++ b/src/linear_algebra/clifford_algebra/equivs.lean @@ -13,6 +13,9 @@ import linear_algebra.quadratic_form.prod /-! # Other constructions isomorphic to Clifford Algebras +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains isomorphisms showing that other types are equivalent to some `clifford_algebra`. ## Rings diff --git a/src/linear_algebra/clifford_algebra/fold.lean b/src/linear_algebra/clifford_algebra/fold.lean index 625219a22945a..353aed894e0f3 100644 --- a/src/linear_algebra/clifford_algebra/fold.lean +++ b/src/linear_algebra/clifford_algebra/fold.lean @@ -8,6 +8,9 @@ import linear_algebra.clifford_algebra.conjugation /-! # Recursive computation rules for the Clifford algebra +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file provides API for a special case `clifford_algebra.foldr` of the universal property `clifford_algebra.lift` with `A = module.End R N` for some arbitrary module `N`. This specialization resembles the `list.foldr` operation, allowing a bilinear map to be "folded" along the generators. diff --git a/src/number_theory/class_number/finite.lean b/src/number_theory/class_number/finite.lean index 3e868fbb2d3b7..bae910d2a816e 100644 --- a/src/number_theory/class_number/finite.lean +++ b/src/number_theory/class_number/finite.lean @@ -14,6 +14,9 @@ import ring_theory.norm /-! # Class numbers of global fields + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. In this file, we use the notion of "admissible absolute value" to prove finiteness of the class group for number fields and function fields, and define `class_number` as the order of this group. diff --git a/src/number_theory/class_number/function_field.lean b/src/number_theory/class_number/function_field.lean index b56a75a38d58a..90b7241ccafc3 100644 --- a/src/number_theory/class_number/function_field.lean +++ b/src/number_theory/class_number/function_field.lean @@ -10,6 +10,9 @@ import number_theory.function_field /-! # Class numbers of function fields +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the class number of a function field as the (finite) cardinality of the class group of its ring of integers. It also proves some elementary results on the class number. diff --git a/src/number_theory/cyclotomic/gal.lean b/src/number_theory/cyclotomic/gal.lean index ee16e36e0619c..168f3db2ae45e 100644 --- a/src/number_theory/cyclotomic/gal.lean +++ b/src/number_theory/cyclotomic/gal.lean @@ -10,6 +10,9 @@ import field_theory.polynomial_galois_group /-! # Galois group of cyclotomic extensions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we show the relationship between the Galois group of `K(ζₙ)` and `(zmod n)ˣ`; it is always a subgroup, and if the `n`th cyclotomic polynomial is irreducible, they are isomorphic. diff --git a/src/number_theory/cyclotomic/rat.lean b/src/number_theory/cyclotomic/rat.lean index 6481a65641071..f285b2f60cf6d 100644 --- a/src/number_theory/cyclotomic/rat.lean +++ b/src/number_theory/cyclotomic/rat.lean @@ -9,6 +9,9 @@ import ring_theory.polynomial.eisenstein.is_integral /-! # Ring of integers of `p ^ n`-th cyclotomic fields + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. We gather results about cyclotomic extensions of `ℚ`. In particular, we compute the ring of integers of a `p ^ n`-th cyclotomic extension of `ℚ`. diff --git a/src/number_theory/number_field/class_number.lean b/src/number_theory/number_field/class_number.lean index 9b1bb64056c76..4958f37408dfe 100644 --- a/src/number_theory/number_field/class_number.lean +++ b/src/number_theory/number_field/class_number.lean @@ -10,6 +10,9 @@ import number_theory.number_field.basic /-! # Class numbers of number fields +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the class number of a number field as the (finite) cardinality of the class group of its ring of integers. It also proves some elementary results on the class number. diff --git a/src/representation_theory/group_cohomology/resolution.lean b/src/representation_theory/group_cohomology/resolution.lean index 434f70e9d0ab5..c31a3397b8fc2 100644 --- a/src/representation_theory/group_cohomology/resolution.lean +++ b/src/representation_theory/group_cohomology/resolution.lean @@ -12,6 +12,9 @@ import representation_theory.Rep /-! # The structure of the `k[G]`-module `k[Gⁿ]` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains facts about an important `k[G]`-module structure on `k[Gⁿ]`, where `k` is a commutative ring and `G` is a group. The module structure arises from the representation `G →* End(k[Gⁿ])` induced by the diagonal action of `G` on `Gⁿ.` diff --git a/src/topology/metric_space/dilation.lean b/src/topology/metric_space/dilation.lean index 24195f12cdfcb..d9b8ecdeb9dc6 100644 --- a/src/topology/metric_space/dilation.lean +++ b/src/topology/metric_space/dilation.lean @@ -10,6 +10,9 @@ import data.fun_like.basic /-! # Dilations +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define dilations, i.e., maps between emetric spaces that satisfy `edist (f x) (f y) = r * edist x y` for some `r ∉ {0, ∞}`. From 3ff3f2d6a3118b8711063de7111a0d77a53219a8 Mon Sep 17 00:00:00 2001 From: Antoine Chambert-Loir Date: Mon, 26 Jun 2023 11:35:20 +0000 Subject: [PATCH 1222/1291] feat(set_theory/cardinal/finite): prove lemmas to handle part_enat.card (#19198) Prove some lemmas that allow to handle part_enat.card of finite cardinals, analogous to those about nat.card Co-authored-by: Antoine Chambert-Loir --- src/data/finite/card.lean | 12 ++++ src/data/nat/part_enat.lean | 6 ++ src/set_theory/cardinal/basic.lean | 81 +++++++++++++++++++++++++- src/set_theory/cardinal/finite.lean | 88 +++++++++++++++++++++++++++++ 4 files changed, 186 insertions(+), 1 deletion(-) diff --git a/src/data/finite/card.lean b/src/data/finite/card.lean index 4191dc9366418..bda3efd6479ac 100644 --- a/src/data/finite/card.lean +++ b/src/data/finite/card.lean @@ -158,6 +158,18 @@ by { haveI := fintype.of_finite α, simpa using fintype.card_subtype_lt hx } end finite +namespace part_enat + +lemma card_eq_coe_nat_card (α : Type*) [finite α] : card α = nat.card α := +begin + unfold part_enat.card, + apply symm, + rw cardinal.coe_nat_eq_to_part_enat_iff, + exact finite.cast_card_eq_mk , +end + +end part_enat + namespace set lemma card_union_le (s t : set α) : nat.card ↥(s ∪ t) ≤ nat.card s + nat.card t := diff --git a/src/data/nat/part_enat.lean b/src/data/nat/part_enat.lean index 7f03e53a943cf..599f44cb6f70e 100644 --- a/src/data/nat/part_enat.lean +++ b/src/data/nat/part_enat.lean @@ -380,6 +380,9 @@ begin apply_mod_cast nat.lt_of_succ_le, apply_mod_cast h end +lemma coe_succ_le_iff {n : ℕ} {e : part_enat} : ↑n.succ ≤ e ↔ ↑n < e:= +by rw [nat.succ_eq_add_one n, nat.cast_add, nat.cast_one, add_one_le_iff_lt (coe_ne_top n)] + lemma lt_add_one_iff_lt {x y : part_enat} (hx : x ≠ ⊤) : x < y + 1 ↔ x ≤ y := begin split, exact le_of_lt_add_one, @@ -388,6 +391,9 @@ begin apply_mod_cast nat.lt_succ_of_le, apply_mod_cast h end +lemma lt_coe_succ_iff_le {x : part_enat} {n : ℕ} (hx : x ≠ ⊤) : x < n.succ ↔ x ≤ n := +by rw [nat.succ_eq_add_one n, nat.cast_add, nat.cast_one, lt_add_one_iff_lt hx] + lemma add_eq_top_iff {a b : part_enat} : a + b = ⊤ ↔ a = ⊤ ∨ b = ⊤ := by apply part_enat.cases_on a; apply part_enat.cases_on b; simp; simp only [(nat.cast_add _ _).symm, part_enat.coe_ne_top]; simp diff --git a/src/set_theory/cardinal/basic.lean b/src/set_theory/cardinal/basic.lean index d3d08052171f7..c04b208884d0f 100644 --- a/src/set_theory/cardinal/basic.lean +++ b/src/set_theory/cardinal/basic.lean @@ -1215,6 +1215,10 @@ by rw [to_nat_apply_of_lt_aleph_0 h, ← classical.some_spec (lt_aleph_0.1 h)] lemma cast_to_nat_of_aleph_0_le {c : cardinal} (h : ℵ₀ ≤ c) : ↑c.to_nat = (0 : cardinal) := by rw [to_nat_apply_of_aleph_0_le h, nat.cast_zero] +lemma to_nat_eq_iff_eq_of_lt_aleph_0 {c d : cardinal} (hc : c < ℵ₀) (hd : d < ℵ₀) : + c.to_nat = d.to_nat ↔ c = d := +by rw [←nat_cast_inj, cast_to_nat_of_lt_aleph_0 hc, cast_to_nat_of_lt_aleph_0 hd] + lemma to_nat_le_iff_le_of_lt_aleph_0 {c d : cardinal} (hc : c < ℵ₀) (hd : d < ℵ₀) : c.to_nat ≤ d.to_nat ↔ c ≤ d := by rw [←nat_cast_le, cast_to_nat_of_lt_aleph_0 hc, cast_to_nat_of_lt_aleph_0 hd] @@ -1357,10 +1361,84 @@ to_part_enat_apply_of_aleph_0_le (infinite_iff.1 h) @[simp] theorem aleph_0_to_part_enat : to_part_enat ℵ₀ = ⊤ := to_part_enat_apply_of_aleph_0_le le_rfl +lemma to_part_enat_eq_top_iff_le_aleph_0 {c : cardinal} : + to_part_enat c = ⊤ ↔ aleph_0 ≤ c := +begin + cases lt_or_ge c aleph_0 with hc hc, + simp only [to_part_enat_apply_of_lt_aleph_0 hc, part_enat.coe_ne_top, false_iff, not_le, hc], + simp only [to_part_enat_apply_of_aleph_0_le hc, eq_self_iff_true, true_iff], + exact hc, +end + +lemma to_part_enat_le_iff_le_of_le_aleph_0 {c c' : cardinal} (h : c ≤ aleph_0) : + to_part_enat c ≤ to_part_enat c' ↔ c ≤ c' := +begin + cases lt_or_ge c aleph_0 with hc hc, + rw to_part_enat_apply_of_lt_aleph_0 hc, + cases lt_or_ge c' aleph_0 with hc' hc', + { rw to_part_enat_apply_of_lt_aleph_0 hc', + rw part_enat.coe_le_coe, + exact to_nat_le_iff_le_of_lt_aleph_0 hc hc', }, + { simp only [to_part_enat_apply_of_aleph_0_le hc', + le_top, true_iff], + exact le_trans h hc', }, + { rw to_part_enat_apply_of_aleph_0_le hc, + simp only [top_le_iff, to_part_enat_eq_top_iff_le_aleph_0, + le_antisymm h hc], }, +end + +lemma to_part_enat_le_iff_le_of_lt_aleph_0 {c c' : cardinal} (hc' : c' < aleph_0) : + to_part_enat c ≤ to_part_enat c' ↔ c ≤ c' := +begin + cases lt_or_ge c aleph_0 with hc hc, + { rw to_part_enat_apply_of_lt_aleph_0 hc, + rw to_part_enat_apply_of_lt_aleph_0 hc', + rw part_enat.coe_le_coe, + exact to_nat_le_iff_le_of_lt_aleph_0 hc hc', }, + { rw to_part_enat_apply_of_aleph_0_le hc, + simp only [top_le_iff, to_part_enat_eq_top_iff_le_aleph_0], + rw [← not_iff_not, not_le, not_le], + simp only [hc', lt_of_lt_of_le hc' hc], }, +end + +lemma to_part_enat_eq_iff_eq_of_le_aleph_0 {c c' : cardinal} + (hc : c ≤ aleph_0) (hc' : c' ≤ aleph_0) : + to_part_enat c = to_part_enat c' ↔ c = c' := by +rw [le_antisymm_iff, le_antisymm_iff, + to_part_enat_le_iff_le_of_le_aleph_0 hc, to_part_enat_le_iff_le_of_le_aleph_0 hc'] + +lemma to_part_enat_mono {c c' : cardinal} (h : c ≤ c') : + to_part_enat c ≤ to_part_enat c' := +begin + cases lt_or_ge c aleph_0 with hc hc, + rw to_part_enat_apply_of_lt_aleph_0 hc, + cases lt_or_ge c' aleph_0 with hc' hc', + rw to_part_enat_apply_of_lt_aleph_0 hc', + simp only [part_enat.coe_le_coe], + exact to_nat_le_of_le_of_lt_aleph_0 hc' h, + rw to_part_enat_apply_of_aleph_0_le hc', + exact le_top, + rw [to_part_enat_apply_of_aleph_0_le hc, + to_part_enat_apply_of_aleph_0_le (le_trans hc h)], +end + lemma to_part_enat_surjective : surjective to_part_enat := λ x, part_enat.cases_on x ⟨ℵ₀, to_part_enat_apply_of_aleph_0_le le_rfl⟩ $ λ n, ⟨n, to_part_enat_cast n⟩ +lemma to_part_enat_lift (c : cardinal.{v}) : (lift.{u v} c).to_part_enat = c.to_part_enat := +begin + cases lt_or_ge c ℵ₀ with hc hc, + { rw [to_part_enat_apply_of_lt_aleph_0 hc, cardinal.to_part_enat_apply_of_lt_aleph_0 _], + simp only [to_nat_lift], + rw [← lift_aleph_0, lift_lt], exact hc }, + { rw [to_part_enat_apply_of_aleph_0_le hc, cardinal.to_part_enat_apply_of_aleph_0_le _], + rw [← lift_aleph_0, lift_le], exact hc } +end + +lemma to_part_enat_congr {β : Type v} (e : α ≃ β) : (#α).to_part_enat = (#β).to_part_enat := +by rw [←to_part_enat_lift, lift_mk_eq.mpr ⟨e⟩, to_part_enat_lift] + lemma mk_to_part_enat_eq_coe_card [fintype α] : (#α).to_part_enat = fintype.card α := by simp @@ -1369,7 +1447,8 @@ lemma mk_int : #ℤ = ℵ₀ := mk_denumerable ℤ lemma mk_pnat : #ℕ+ = ℵ₀ := mk_denumerable ℕ+ /-- **König's theorem** -/ -theorem sum_lt_prod {ι} (f g : ι → cardinal) (H : ∀ i, f i < g i) : sum f < prod g := +theorem sum_lt_prod {ι} (f g : ι → cardinal) (H : ∀ i, f i < g i) : +sum f < prod g := lt_of_not_ge $ λ ⟨F⟩, begin haveI : inhabited (Π (i : ι), (g i).out), { refine ⟨λ i, classical.choice $ mk_ne_zero_iff.1 _⟩, diff --git a/src/set_theory/cardinal/finite.lean b/src/set_theory/cardinal/finite.lean index 5495b095f9091..9e98d2bb67b27 100644 --- a/src/set_theory/cardinal/finite.lean +++ b/src/set_theory/cardinal/finite.lean @@ -120,4 +120,92 @@ lemma card_eq_coe_fintype_card [fintype α] : card α = fintype.card α := mk_to @[simp] lemma card_eq_top_of_infinite [infinite α] : card α = ⊤ := mk_to_part_enat_of_infinite +lemma card_congr {α : Type*} {β : Type*} (f : α ≃ β) : + part_enat.card α = part_enat.card β := +cardinal.to_part_enat_congr f + +lemma card_ulift (α : Type*) : card (ulift α) = card α := +card_congr equiv.ulift + +@[simp] lemma card_plift (α : Type*) : card (plift α) = card α := +card_congr equiv.plift + +lemma card_image_of_inj_on {α : Type*} {β : Type*} {f : α → β} {s : set α} (h : set.inj_on f s) : + card (f '' s) = card s := +card_congr (equiv.set.image_of_inj_on f s h).symm + +lemma card_image_of_injective {α : Type*} {β : Type*} + (f : α → β) (s : set α) (h : function.injective f) : + card (f '' s) = card s := +card_image_of_inj_on (set.inj_on_of_injective h s) + +-- Should I keep the 6 following lemmas ? +@[simp] +lemma _root_.cardinal.coe_nat_le_to_part_enat_iff {n : ℕ} {c : cardinal} : + ↑n ≤ to_part_enat c ↔ ↑n ≤ c := +by rw [← to_part_enat_cast n, to_part_enat_le_iff_le_of_le_aleph_0 (le_of_lt (nat_lt_aleph_0 n))] + +@[simp] +lemma _root_.cardinal.to_part_enat_le_coe_nat_iff {c : cardinal} {n : ℕ} : + to_part_enat c ≤ n ↔ c ≤ n := +by rw [← to_part_enat_cast n, + to_part_enat_le_iff_le_of_lt_aleph_0 (nat_lt_aleph_0 n)] + +@[simp] +lemma _root_.cardinal.coe_nat_eq_to_part_enat_iff {n : ℕ} {c : cardinal} : + ↑n = to_part_enat c ↔ ↑n = c := +by rw [le_antisymm_iff, le_antisymm_iff, + cardinal.coe_nat_le_to_part_enat_iff, cardinal.to_part_enat_le_coe_nat_iff] + +@[simp] +lemma _root_.cardinal.to_part_enat_eq_coe_nat_iff {c : cardinal} {n : ℕ} : + to_part_enat c = n ↔ c = n:= +by rw [eq_comm, cardinal.coe_nat_eq_to_part_enat_iff, eq_comm] + +@[simp] +lemma _root_.cardinal.coe_nat_lt_coe_iff_lt {n : ℕ} {c : cardinal} : + ↑n < to_part_enat c ↔ ↑n < c := +by simp only [← not_le, cardinal.to_part_enat_le_coe_nat_iff] + +@[simp] +lemma _root_.cardinal.lt_coe_nat_iff_lt {n : ℕ} {c : cardinal} : + to_part_enat c < n ↔ c < n := +by simp only [← not_le, cardinal.coe_nat_le_to_part_enat_iff] + +lemma card_eq_zero_iff_empty (α : Type*) : card α = 0 ↔ is_empty α := +begin + rw ← cardinal.mk_eq_zero_iff, + conv_rhs { rw ← nat.cast_zero }, + rw ← cardinal.to_part_enat_eq_coe_nat_iff, + simp only [part_enat.card, nat.cast_zero] +end + +lemma card_le_one_iff_subsingleton (α : Type*) : card α ≤ 1 ↔ subsingleton α := +begin + rw ← le_one_iff_subsingleton, + conv_rhs { rw ← nat.cast_one}, + rw ← cardinal.to_part_enat_le_coe_nat_iff, + simp only [part_enat.card, nat.cast_one] +end + +lemma one_lt_card_iff_nontrivial (α : Type*) : 1 < card α ↔ nontrivial α := +begin + rw ← one_lt_iff_nontrivial, + conv_rhs { rw ← nat.cast_one}, + rw ← cardinal.coe_nat_lt_coe_iff_lt, + simp only [part_enat.card, nat.cast_one] +end + +lemma is_finite_of_card {α : Type*} {n : ℕ} (hα : part_enat.card α = n) : + finite α := +begin + apply or.resolve_right (finite_or_infinite α), + intro h, resetI, + apply part_enat.coe_ne_top n, + rw ← hα, + exact part_enat.card_eq_top_of_infinite, +end + + + end part_enat From 8b981918a93bc45a8600de608cde7944a80d92b9 Mon Sep 17 00:00:00 2001 From: Moritz Doll Date: Mon, 26 Jun 2023 15:07:59 +0000 Subject: [PATCH 1223/1291] feat(analysis/inner_product_space): the adjoint for unbounded operators (#18820) --- docs/overview.yaml | 1 + .../inner_product_space/linear_pmap.lean | 212 ++++++++++++++++++ src/linear_algebra/linear_pmap.lean | 2 + 3 files changed, 215 insertions(+) create mode 100644 src/analysis/inner_product_space/linear_pmap.lean diff --git a/docs/overview.yaml b/docs/overview.yaml index d1653f00c6261..ecc7271289e99 100644 --- a/docs/overview.yaml +++ b/docs/overview.yaml @@ -274,6 +274,7 @@ Analysis: Hilbert spaces: Inner product space, over $R$ or $C$: 'inner_product_space' Cauchy-Schwarz inequality: 'inner_mul_inner_self_le' + adjoint operator: 'linear_pmap.adjoint' self-adjoint operator: 'is_self_adjoint' orthogonal projection: 'orthogonal_projection' reflection: 'reflection' diff --git a/src/analysis/inner_product_space/linear_pmap.lean b/src/analysis/inner_product_space/linear_pmap.lean new file mode 100644 index 0000000000000..8d36210252e4c --- /dev/null +++ b/src/analysis/inner_product_space/linear_pmap.lean @@ -0,0 +1,212 @@ +/- +Copyright (c) 2022 Moritz Doll. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Moritz Doll +-/ + +import analysis.inner_product_space.adjoint +import topology.algebra.module.linear_pmap +import topology.algebra.module.basic + +/-! + +# Partially defined linear operators on Hilbert spaces + +We will develop the basics of the theory of unbounded operators on Hilbert spaces. + +## Main definitions + +* `linear_pmap.is_formal_adjoint`: An operator `T` is a formal adjoint of `S` if for all `x` in the + domain of `T` and `y` in the domain of `S`, we have that `⟪T x, y⟫ = ⟪x, S y⟫`. +* `linear_pmap.adjoint`: The adjoint of a map `E →ₗ.[𝕜] F` as a map `F →ₗ.[𝕜] E`. + +## Main statements + +* `linear_pmap.adjoint_is_formal_adjoint`: The adjoint is a formal adjoint +* `linear_pmap.is_formal_adjoint.le_adjoint`: Every formal adjoint is contained in the adjoint +* `continuous_linear_map.to_pmap_adjoint_eq_adjoint_to_pmap_of_dense`: The adjoint on + `continuous_linear_map` and `linear_pmap` coincide. + +## Notation + +* For `T : E →ₗ.[𝕜] F` the adjoint can be written as `T†`. + This notation is localized in `linear_pmap`. + +## Implementation notes + +We use the junk value pattern to define the adjoint for all `linear_pmap`s. In the case that +`T : E →ₗ.[𝕜] F` is not densely defined the adjoint `T†` is the zero map from `T.adjoint_domain` to +`E`. + +## References + +* [J. Weidmann, *Linear Operators in Hilbert Spaces*][weidmann_linear] + +## Tags + +Unbounded operators, closed operators +-/ + + +noncomputable theory + +open is_R_or_C +open_locale complex_conjugate classical + +variables {𝕜 E F G : Type*} [is_R_or_C 𝕜] +variables [normed_add_comm_group E] [inner_product_space 𝕜 E] +variables [normed_add_comm_group F] [inner_product_space 𝕜 F] + +local notation `⟪`x`, `y`⟫` := @inner 𝕜 _ _ x y + +namespace linear_pmap + +/-- An operator `T` is a formal adjoint of `S` if for all `x` in the domain of `T` and `y` in the +domain of `S`, we have that `⟪T x, y⟫ = ⟪x, S y⟫`. -/ +def is_formal_adjoint (T : E →ₗ.[𝕜] F) (S : F →ₗ.[𝕜] E) : Prop := +∀ (x : T.domain) (y : S.domain), ⟪T x, y⟫ = ⟪(x : E), S y⟫ + +variables {T : E →ₗ.[𝕜] F} {S : F →ₗ.[𝕜] E} + +@[protected] lemma is_formal_adjoint.symm (h : T.is_formal_adjoint S) : S.is_formal_adjoint T := +λ y _, by rw [←inner_conj_symm, ←inner_conj_symm (y : F), h] + +variables (T) + +/-- The domain of the adjoint operator. + +This definition is needed to construct the adjoint operator and the preferred version to use is +`T.adjoint.domain` instead of `T.adjoint_domain`. -/ +def adjoint_domain : submodule 𝕜 F := +{ carrier := {y | continuous ((innerₛₗ 𝕜 y).comp T.to_fun)}, + zero_mem' := by { rw [set.mem_set_of_eq, linear_map.map_zero, linear_map.zero_comp], + exact continuous_zero }, + add_mem' := λ x y hx hy, by { rw [set.mem_set_of_eq, linear_map.map_add] at *, exact hx.add hy }, + smul_mem' := λ a x hx, by { rw [set.mem_set_of_eq, linear_map.map_smulₛₗ] at *, + exact hx.const_smul (conj a) } } + +/-- The operator `λ x, ⟪y, T x⟫` considered as a continuous linear operator from `T.adjoint_domain` +to `𝕜`. -/ +def adjoint_domain_mk_clm (y : T.adjoint_domain) : T.domain →L[𝕜] 𝕜 := +⟨(innerₛₗ 𝕜 (y : F)).comp T.to_fun, y.prop⟩ + +lemma adjoint_domain_mk_clm_apply (y : T.adjoint_domain) (x : T.domain) : + adjoint_domain_mk_clm T y x = ⟪(y : F), T x⟫ := rfl + +variable {T} +variable (hT : dense (T.domain : set E)) + +include hT + +/-- The unique continuous extension of the operator `adjoint_domain_mk_clm` to `E`. -/ +def adjoint_domain_mk_clm_extend (y : T.adjoint_domain) : + E →L[𝕜] 𝕜 := +(T.adjoint_domain_mk_clm y).extend (submodule.subtypeL T.domain) + hT.dense_range_coe uniform_embedding_subtype_coe.to_uniform_inducing + +@[simp] lemma adjoint_domain_mk_clm_extend_apply (y : T.adjoint_domain) (x : T.domain) : + adjoint_domain_mk_clm_extend hT y (x : E) = ⟪(y : F), T x⟫ := +continuous_linear_map.extend_eq _ _ _ _ _ + +variables [complete_space E] + +/-- The adjoint as a linear map from its domain to `E`. + +This is an auxiliary definition needed to define the adjoint operator as a `linear_pmap` without +the assumption that `T.domain` is dense. -/ +def adjoint_aux : T.adjoint_domain →ₗ[𝕜] E := +{ to_fun := λ y, (inner_product_space.to_dual 𝕜 E).symm (adjoint_domain_mk_clm_extend hT y), + map_add' := λ x y, hT.eq_of_inner_left $ λ _, + by simp only [inner_add_left, submodule.coe_add, inner_product_space.to_dual_symm_apply, + adjoint_domain_mk_clm_extend_apply], + map_smul' := λ _ _, hT.eq_of_inner_left $ λ _, + by simp only [inner_smul_left, submodule.coe_smul_of_tower, ring_hom.id_apply, + inner_product_space.to_dual_symm_apply, adjoint_domain_mk_clm_extend_apply] } + +lemma adjoint_aux_inner (y : T.adjoint_domain) (x : T.domain) : + ⟪adjoint_aux hT y, x⟫ = ⟪(y : F), T x⟫ := +by simp only [adjoint_aux, linear_map.coe_mk, inner_product_space.to_dual_symm_apply, + adjoint_domain_mk_clm_extend_apply] + +lemma adjoint_aux_unique (y : T.adjoint_domain) {x₀ : E} + (hx₀ : ∀ x : T.domain, ⟪x₀, x⟫ = ⟪(y : F), T x⟫) : adjoint_aux hT y = x₀ := +hT.eq_of_inner_left (λ v, (adjoint_aux_inner hT _ _).trans (hx₀ v).symm) + +omit hT + +variable (T) + +/-- The adjoint operator as a partially defined linear operator. -/ +def adjoint : F →ₗ.[𝕜] E := +{ domain := T.adjoint_domain, + to_fun := if hT : dense (T.domain : set E) then adjoint_aux hT else 0 } + +localized "postfix (name := adjoint) `†`:1100 := linear_pmap.adjoint" in linear_pmap + +lemma mem_adjoint_domain_iff (y : F) : + y ∈ T†.domain ↔ continuous ((innerₛₗ 𝕜 y).comp T.to_fun) := iff.rfl + +variable {T} + +lemma mem_adjoint_domain_of_exists (y : F) (h : ∃ w : E, ∀ (x : T.domain), ⟪w, x⟫ = ⟪y, T x⟫) : + y ∈ T†.domain := +begin + cases h with w hw, + rw T.mem_adjoint_domain_iff, + have : continuous ((innerSL 𝕜 w).comp T.domain.subtypeL) := by continuity, + convert this using 1, + exact funext (λ x, (hw x).symm), +end + +lemma adjoint_apply_of_not_dense (hT : ¬ dense (T.domain : set E)) (y : T†.domain) : T† y = 0 := +begin + change (if hT : dense (T.domain : set E) then adjoint_aux hT else 0) y = _, + simp only [hT, not_false_iff, dif_neg, linear_map.zero_apply], +end + +include hT + +lemma adjoint_apply_of_dense (y : T†.domain) : T† y = adjoint_aux hT y := +begin + change (if hT : dense (T.domain : set E) then adjoint_aux hT else 0) y = _, + simp only [hT, dif_pos, linear_map.coe_mk], +end + +lemma adjoint_apply_eq (y : T†.domain) {x₀ : E} + (hx₀ : ∀ x : T.domain, ⟪x₀, x⟫ = ⟪(y : F), T x⟫) : T† y = x₀ := +(adjoint_apply_of_dense hT y).symm ▸ adjoint_aux_unique hT _ hx₀ + +/-- The fundamental property of the adjoint. -/ +lemma adjoint_is_formal_adjoint : T†.is_formal_adjoint T := +λ x, (adjoint_apply_of_dense hT x).symm ▸ adjoint_aux_inner hT x + +/-- The adjoint is maximal in the sense that it contains every formal adjoint. -/ +lemma is_formal_adjoint.le_adjoint (h : T.is_formal_adjoint S) : S ≤ T† := +-- Trivially, every `x : S.domain` is in `T.adjoint.domain` +⟨λ x hx, mem_adjoint_domain_of_exists _ ⟨S ⟨x, hx⟩, h.symm ⟨x, hx⟩⟩, + -- Equality on `S.domain` follows from equality + -- `⟪v, S x⟫ = ⟪v, T.adjoint y⟫` for all `v : T.domain`: + λ _ _ hxy, (adjoint_apply_eq hT _ (λ _, by rw [h.symm, hxy])).symm⟩ + +end linear_pmap + +namespace continuous_linear_map + +variables [complete_space E] [complete_space F] +variables (A : E →L[𝕜] F) {p : submodule 𝕜 E} + +/-- Restricting `A` to a dense submodule and taking the `linear_pmap.adjoint` is the same +as taking the `continuous_linear_map.adjoint` interpreted as a `linear_pmap`. -/ +lemma to_pmap_adjoint_eq_adjoint_to_pmap_of_dense (hp : dense (p : set E)) : + (A.to_pmap p).adjoint = A.adjoint.to_pmap ⊤ := +begin + ext, + { simp only [to_linear_map_eq_coe, linear_map.to_pmap_domain, submodule.mem_top, iff_true, + linear_pmap.mem_adjoint_domain_iff, linear_map.coe_comp, innerₛₗ_apply_coe], + exact ((innerSL 𝕜 x).comp $ A.comp $ submodule.subtypeL _).cont }, + intros x y hxy, + refine linear_pmap.adjoint_apply_eq hp _ (λ v, _), + simp only [adjoint_inner_left, hxy, linear_map.to_pmap_apply, to_linear_map_eq_coe, coe_coe], +end + +end continuous_linear_map diff --git a/src/linear_algebra/linear_pmap.lean b/src/linear_algebra/linear_pmap.lean index 17569ffbfd37a..a20d1d6a41afb 100644 --- a/src/linear_algebra/linear_pmap.lean +++ b/src/linear_algebra/linear_pmap.lean @@ -470,6 +470,8 @@ def to_pmap (f : E →ₗ[R] F) (p : submodule R E) : E →ₗ.[R] F := @[simp] lemma to_pmap_apply (f : E →ₗ[R] F) (p : submodule R E) (x : p) : f.to_pmap p x = f x := rfl +@[simp] lemma to_pmap_domain (f : E →ₗ[R] F) (p : submodule R E) : (f.to_pmap p).domain = p := rfl + /-- Compose a linear map with a `linear_pmap` -/ def comp_pmap (g : F →ₗ[R] G) (f : E →ₗ.[R] F) : E →ₗ.[R] G := { domain := f.domain, From 15db1b4f26ba89c6eb0c78b0a44c7e779a788e29 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Tue, 27 Jun 2023 04:28:35 +0000 Subject: [PATCH 1224/1291] chore(category_theory/limits/construction/over): rename default to basic (#19217) This was a default file with content, so needs renaming. Afterwards it can just be ported as usual. Co-authored-by: Scott Morrison --- .../limits/constructions/over/{default.lean => basic.lean} | 0 1 file changed, 0 insertions(+), 0 deletions(-) rename src/category_theory/limits/constructions/over/{default.lean => basic.lean} (100%) diff --git a/src/category_theory/limits/constructions/over/default.lean b/src/category_theory/limits/constructions/over/basic.lean similarity index 100% rename from src/category_theory/limits/constructions/over/default.lean rename to src/category_theory/limits/constructions/over/basic.lean From 08b63ab58a6ec1157ebeafcbbe6c7a3fb3c9f6d5 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Wed, 28 Jun 2023 06:08:04 +0000 Subject: [PATCH 1225/1291] chore(*): add mathlib4 synchronization comments (#19216) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.category.Mon.colimits` * `algebra.category.fgModule.basic` * `analysis.complex.upper_half_plane.basic` * `analysis.inner_product_space.linear_pmap` * `analysis.special_functions.gamma.beta` * `analysis.special_functions.gamma.bohr_mollerup` * `analysis.special_functions.gaussian` * `category_theory.closed.ideal` * `category_theory.monoidal.Bimod` * `combinatorics.simple_graph.regularity.bound` * `combinatorics.simple_graph.regularity.chunk` * `combinatorics.simple_graph.regularity.increment` * `combinatorics.simple_graph.regularity.lemma` * `control.random` * `geometry.manifold.complex` * `geometry.manifold.mfderiv` * `geometry.manifold.partition_of_unity` * `linear_algebra.exterior_algebra.of_alternating` * `logic.equiv.array` * `number_theory.bertrand` * `number_theory.legendre_symbol.gauss_eisenstein_lemmas` * `number_theory.legendre_symbol.gauss_sum` * `number_theory.legendre_symbol.jacobi_symbol` * `number_theory.legendre_symbol.quadratic_char.gauss_sum` * `number_theory.legendre_symbol.quadratic_reciprocity` * `number_theory.sum_two_squares` * `number_theory.zsqrtd.quadratic_reciprocity` * `ring_theory.dedekind_domain.selmer_group` * `ring_theory.witt_vector.frobenius` * `ring_theory.witt_vector.identities` * `ring_theory.witt_vector.init_tail` * `ring_theory.witt_vector.verschiebung` * `set_theory.game.basic` * `set_theory.game.impartial` * `set_theory.game.ordinal` * `set_theory.surreal.basic` * `testing.slim_check.gen` * `testing.slim_check.sampleable` * `testing.slim_check.testable` --- src/algebra/category/Mon/colimits.lean | 3 +++ src/algebra/category/fgModule/basic.lean | 3 +++ src/analysis/special_functions/gamma/bohr_mollerup.lean | 3 +++ src/analysis/special_functions/gaussian.lean | 3 +++ src/category_theory/monoidal/Bimod.lean | 3 +++ src/combinatorics/simple_graph/regularity/bound.lean | 3 +++ src/combinatorics/simple_graph/regularity/chunk.lean | 3 +++ src/combinatorics/simple_graph/regularity/increment.lean | 3 +++ src/combinatorics/simple_graph/regularity/lemma.lean | 3 +++ src/geometry/manifold/mfderiv.lean | 3 +++ src/linear_algebra/exterior_algebra/of_alternating.lean | 3 +++ src/number_theory/bertrand.lean | 3 +++ src/number_theory/legendre_symbol/gauss_eisenstein_lemmas.lean | 3 +++ src/number_theory/legendre_symbol/gauss_sum.lean | 3 +++ .../legendre_symbol/quadratic_char/gauss_sum.lean | 3 +++ src/number_theory/legendre_symbol/quadratic_reciprocity.lean | 3 +++ src/number_theory/sum_two_squares.lean | 3 +++ src/number_theory/zsqrtd/quadratic_reciprocity.lean | 3 +++ src/ring_theory/dedekind_domain/selmer_group.lean | 3 +++ src/set_theory/game/basic.lean | 3 +++ src/set_theory/game/ordinal.lean | 3 +++ 21 files changed, 63 insertions(+) diff --git a/src/algebra/category/Mon/colimits.lean b/src/algebra/category/Mon/colimits.lean index 9f12277c95cb3..b1bbcb32c2f8d 100644 --- a/src/algebra/category/Mon/colimits.lean +++ b/src/algebra/category/Mon/colimits.lean @@ -10,6 +10,9 @@ import category_theory.concrete_category.elementwise /-! # The category of monoids has all colimits. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We do this construction knowing nothing about monoids. In particular, I want to claim that this file could be produced by a python script that just looks at the output of `#print monoid`: diff --git a/src/algebra/category/fgModule/basic.lean b/src/algebra/category/fgModule/basic.lean index f70d4ec8e3ecb..118a4688f6a66 100644 --- a/src/algebra/category/fgModule/basic.lean +++ b/src/algebra/category/fgModule/basic.lean @@ -12,6 +12,9 @@ import algebra.category.Module.monoidal.closed /-! # The category of finitely generated modules over a ring +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This introduces `fgModule R`, the category of finitely generated modules over a ring `R`. It is implemented as a full subcategory on a subtype of `Module R`. diff --git a/src/analysis/special_functions/gamma/bohr_mollerup.lean b/src/analysis/special_functions/gamma/bohr_mollerup.lean index 9db12f986223c..3a6b66bd0659d 100644 --- a/src/analysis/special_functions/gamma/bohr_mollerup.lean +++ b/src/analysis/special_functions/gamma/bohr_mollerup.lean @@ -8,6 +8,9 @@ import analysis.special_functions.gaussian /-! # Convexity properties of the Gamma function +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we prove that `Gamma` and `log ∘ Gamma` are convex functions on the positive real line. We then prove the Bohr-Mollerup theorem, which characterises `Gamma` as the *unique* positive-real-valued, log-convex function on the positive reals satisfying `f (x + 1) = x f x` and diff --git a/src/analysis/special_functions/gaussian.lean b/src/analysis/special_functions/gaussian.lean index b21f166f048e6..7705bdc538282 100644 --- a/src/analysis/special_functions/gaussian.lean +++ b/src/analysis/special_functions/gaussian.lean @@ -13,6 +13,9 @@ import analysis.fourier.poisson_summation /-! # Gaussian integral +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We prove various versions of the formula for the Gaussian integral: * `integral_gaussian`: for real `b` we have `∫ x:ℝ, exp (-b * x^2) = sqrt (π / b)`. * `integral_gaussian_complex`: for complex `b` with `0 < re b` we have diff --git a/src/category_theory/monoidal/Bimod.lean b/src/category_theory/monoidal/Bimod.lean index 7ff041dcfd99e..48b955f7b99b4 100644 --- a/src/category_theory/monoidal/Bimod.lean +++ b/src/category_theory/monoidal/Bimod.lean @@ -9,6 +9,9 @@ import category_theory.limits.preserves.shapes.equalizers /-! # The category of bimodule objects over a pair of monoid objects. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ universes v₁ v₂ u₁ u₂ diff --git a/src/combinatorics/simple_graph/regularity/bound.lean b/src/combinatorics/simple_graph/regularity/bound.lean index 2241cf244ebaf..6e2e8992eac1a 100644 --- a/src/combinatorics/simple_graph/regularity/bound.lean +++ b/src/combinatorics/simple_graph/regularity/bound.lean @@ -10,6 +10,9 @@ import order.partition.equipartition /-! # Numerical bounds for Szemerédi Regularity Lemma +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file gathers the numerical facts required by the proof of Szemerédi's regularity lemma. This entire file is internal to the proof of Szemerédi Regularity Lemma. diff --git a/src/combinatorics/simple_graph/regularity/chunk.lean b/src/combinatorics/simple_graph/regularity/chunk.lean index ab766c66b386c..a72af72eec1e4 100644 --- a/src/combinatorics/simple_graph/regularity/chunk.lean +++ b/src/combinatorics/simple_graph/regularity/chunk.lean @@ -10,6 +10,9 @@ import combinatorics.simple_graph.regularity.uniform /-! # Chunk of the increment partition for Szemerédi Regularity Lemma +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In the proof of Szemerédi Regularity Lemma, we need to partition each part of a starting partition to increase the energy. This file defines those partitions of parts and shows that they locally increase the energy. diff --git a/src/combinatorics/simple_graph/regularity/increment.lean b/src/combinatorics/simple_graph/regularity/increment.lean index 8f0c29a8e2a28..b19bae79f86d9 100644 --- a/src/combinatorics/simple_graph/regularity/increment.lean +++ b/src/combinatorics/simple_graph/regularity/increment.lean @@ -9,6 +9,9 @@ import combinatorics.simple_graph.regularity.energy /-! # Increment partition for Szemerédi Regularity Lemma +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In the proof of Szemerédi Regularity Lemma, we need to partition each part of a starting partition to increase the energy. This file defines the partition obtained by gluing the parts partitions together (the *increment partition*) and shows that the energy globally increases. diff --git a/src/combinatorics/simple_graph/regularity/lemma.lean b/src/combinatorics/simple_graph/regularity/lemma.lean index 0b271479520e1..88ae521886eaa 100644 --- a/src/combinatorics/simple_graph/regularity/lemma.lean +++ b/src/combinatorics/simple_graph/regularity/lemma.lean @@ -8,6 +8,9 @@ import combinatorics.simple_graph.regularity.increment /-! # Szemerédi's Regularity Lemma +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we prove Szemerédi's Regularity Lemma (aka SRL). This is a landmark result in combinatorics roughly stating that any sufficiently big graph behaves like a random graph. This is useful because random graphs are well-behaved in many aspects. diff --git a/src/geometry/manifold/mfderiv.lean b/src/geometry/manifold/mfderiv.lean index 67669519a2b97..30c9eb15fbf17 100644 --- a/src/geometry/manifold/mfderiv.lean +++ b/src/geometry/manifold/mfderiv.lean @@ -8,6 +8,9 @@ import geometry.manifold.vector_bundle.tangent /-! # The derivative of functions between smooth manifolds +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Let `M` and `M'` be two smooth manifolds with corners over a field `𝕜` (with respective models with corners `I` on `(E, H)` and `I'` on `(E', H')`), and let `f : M → M'`. We define the derivative of the function at a point, within a set or along the whole space, mimicking the API diff --git a/src/linear_algebra/exterior_algebra/of_alternating.lean b/src/linear_algebra/exterior_algebra/of_alternating.lean index 0aaeb47af7172..ad0ae0b28a5be 100644 --- a/src/linear_algebra/exterior_algebra/of_alternating.lean +++ b/src/linear_algebra/exterior_algebra/of_alternating.lean @@ -10,6 +10,9 @@ import linear_algebra.exterior_algebra.basic /-! # Extending an alternating map to the exterior algebra +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions * `exterior_algebra.lift_alternating`: construct a linear map out of the exterior algebra diff --git a/src/number_theory/bertrand.lean b/src/number_theory/bertrand.lean index 38e6d956438a7..73927a0e9b79c 100644 --- a/src/number_theory/bertrand.lean +++ b/src/number_theory/bertrand.lean @@ -12,6 +12,9 @@ import analysis.convex.specific_functions.deriv /-! # Bertrand's Postulate +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains a proof of Bertrand's postulate: That between any positive number and its double there is a prime. diff --git a/src/number_theory/legendre_symbol/gauss_eisenstein_lemmas.lean b/src/number_theory/legendre_symbol/gauss_eisenstein_lemmas.lean index 75380ab52d02e..344697e529726 100644 --- a/src/number_theory/legendre_symbol/gauss_eisenstein_lemmas.lean +++ b/src/number_theory/legendre_symbol/gauss_eisenstein_lemmas.lean @@ -8,6 +8,9 @@ import number_theory.legendre_symbol.quadratic_reciprocity /-! # Lemmas of Gauss and Eisenstein +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains the Lemmas of Gauss and Eisenstein on the Legendre symbol. The main results are `zmod.gauss_lemma` and `zmod.eisenstein_lemma`. -/ diff --git a/src/number_theory/legendre_symbol/gauss_sum.lean b/src/number_theory/legendre_symbol/gauss_sum.lean index 43f375c18d498..d3eec3b506a90 100644 --- a/src/number_theory/legendre_symbol/gauss_sum.lean +++ b/src/number_theory/legendre_symbol/gauss_sum.lean @@ -10,6 +10,9 @@ import algebra.char_p.char_and_card /-! # Gauss sums +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define the Gauss sum associated to a multiplicative and an additive character of a finite field and prove some results about them. diff --git a/src/number_theory/legendre_symbol/quadratic_char/gauss_sum.lean b/src/number_theory/legendre_symbol/quadratic_char/gauss_sum.lean index 6e220b86a5d81..ef6cb184b7bcd 100644 --- a/src/number_theory/legendre_symbol/quadratic_char/gauss_sum.lean +++ b/src/number_theory/legendre_symbol/quadratic_char/gauss_sum.lean @@ -9,6 +9,9 @@ import number_theory.legendre_symbol.gauss_sum /-! # Quadratic characters of finite fields +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Further facts relying on Gauss sums. -/ diff --git a/src/number_theory/legendre_symbol/quadratic_reciprocity.lean b/src/number_theory/legendre_symbol/quadratic_reciprocity.lean index 53cfb747b8efe..c8b13ea94189e 100644 --- a/src/number_theory/legendre_symbol/quadratic_reciprocity.lean +++ b/src/number_theory/legendre_symbol/quadratic_reciprocity.lean @@ -9,6 +9,9 @@ import number_theory.legendre_symbol.quadratic_char.gauss_sum /-! # Quadratic reciprocity. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main results We prove the law of quadratic reciprocity, see `legendre_sym.quadratic_reciprocity` and diff --git a/src/number_theory/sum_two_squares.lean b/src/number_theory/sum_two_squares.lean index fe7bbf6eae2f2..863876134230f 100644 --- a/src/number_theory/sum_two_squares.lean +++ b/src/number_theory/sum_two_squares.lean @@ -10,6 +10,9 @@ import tactic.linear_combination /-! # Sums of two squares +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Fermat's theorem on the sum of two squares. Every prime `p` congruent to 1 mod 4 is the sum of two squares; see `nat.prime.sq_add_sq` (which has the weaker assumption `p % 4 ≠ 3`). diff --git a/src/number_theory/zsqrtd/quadratic_reciprocity.lean b/src/number_theory/zsqrtd/quadratic_reciprocity.lean index 876937bec856e..bb4db5899c3c4 100644 --- a/src/number_theory/zsqrtd/quadratic_reciprocity.lean +++ b/src/number_theory/zsqrtd/quadratic_reciprocity.lean @@ -9,6 +9,9 @@ import number_theory.legendre_symbol.quadratic_reciprocity /-! # Facts about the gaussian integers relying on quadratic reciprocity. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main statements `prime_iff_mod_four_eq_three_of_nat_prime` diff --git a/src/ring_theory/dedekind_domain/selmer_group.lean b/src/ring_theory/dedekind_domain/selmer_group.lean index f729e75d1c898..adcbbfc9e1fb3 100644 --- a/src/ring_theory/dedekind_domain/selmer_group.lean +++ b/src/ring_theory/dedekind_domain/selmer_group.lean @@ -11,6 +11,9 @@ import ring_theory.norm /-! # Selmer groups of fraction fields of Dedekind domains +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Let $K$ be the field of fractions of a Dedekind domain $R$. For any set $S$ of prime ideals in the height one spectrum of $R$, and for any natural number $n$, the Selmer group $K(S, n)$ is defined to be the subgroup of the unit group $K^\times$ modulo $n$-th powers where each element has $v$-adic diff --git a/src/set_theory/game/basic.lean b/src/set_theory/game/basic.lean index c440ab66e49f4..e98541fca9c39 100644 --- a/src/set_theory/game/basic.lean +++ b/src/set_theory/game/basic.lean @@ -9,6 +9,9 @@ import tactic.abel /-! # Combinatorial games. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define the quotient of pre-games by the equivalence relation `p ≈ q ↔ p ≤ q ∧ q ≤ p` (its `antisymmetrization`), and construct an instance `add_comm_group game`, as well as an instance `partial_order game`. diff --git a/src/set_theory/game/ordinal.lean b/src/set_theory/game/ordinal.lean index 34baa6c2abfb3..1fb9ef66b97e6 100644 --- a/src/set_theory/game/ordinal.lean +++ b/src/set_theory/game/ordinal.lean @@ -10,6 +10,9 @@ import set_theory.ordinal.natural_ops /-! # Ordinals as games +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define the canonical map `ordinal → pgame`, where every ordinal is mapped to the game whose left set consists of all previous ordinals. From 147b294346843885f952c5171e9606616a8fd869 Mon Sep 17 00:00:00 2001 From: Apurva Date: Thu, 29 Jun 2023 07:01:52 +0000 Subject: [PATCH 1226/1291] feat(analysis/convex/cone/proper): add hyperplane_separation and comap (#19008) We add the theorem `hyperplane_separation` which is a relative version of [convex_cone.hyperplane_separation_of_nonempty_of_is_closed_of_nmem](https://leanprover-community.github.io/mathlib_docs/analysis/convex/cone/dual.html#convex_cone.hyperplane_separation_of_nonempty_of_is_closed_of_nmem). This is the most general form of Farkas' lemma (that I know of) for convex cones. We also add `proper_cone.comap` and a few theorems about it. Co-authored-by: Apurva Nakade --- src/analysis/convex/cone/proper.lean | 82 ++++++++++++++++++++++++++-- 1 file changed, 78 insertions(+), 4 deletions(-) diff --git a/src/analysis/convex/cone/proper.lean b/src/analysis/convex/cone/proper.lean index e85a633104328..9babbda4c6ef3 100644 --- a/src/analysis/convex/cone/proper.lean +++ b/src/analysis/convex/cone/proper.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Apurva Nakade -/ import analysis.convex.cone.dual +import analysis.inner_product_space.adjoint /-! # Proper cones @@ -17,8 +18,6 @@ linear programs, the results from this file can be used to prove duality theorem ## TODO The next steps are: -- Prove the cone version of Farkas' lemma (2.3.4 in the reference). -- Add comap, adjoint - Add convex_cone_class that extends set_like and replace the below instance - Define the positive cone as a proper cone. - Define primal and dual cone programs and prove weak duality. @@ -33,6 +32,8 @@ The next steps are: -/ +open continuous_linear_map filter set + namespace convex_cone variables {𝕜 : Type*} [ordered_semiring 𝕜] @@ -118,12 +119,13 @@ section inner_product_space variables {E : Type*} [normed_add_comm_group E] [inner_product_space ℝ E] variables {F : Type*} [normed_add_comm_group F] [inner_product_space ℝ F] +variables {G : Type*} [normed_add_comm_group G] [inner_product_space ℝ G] protected lemma pointed (K : proper_cone ℝ E) : (K : convex_cone ℝ E).pointed := (K : convex_cone ℝ E).pointed_of_nonempty_of_is_closed K.nonempty K.is_closed /-- The closure of image of a proper cone under a continuous `ℝ`-linear map is a proper cone. We -use continuous maps here so that the adjoint of f is also a map between proper cones. -/ +use continuous maps here so that the comap of f is also a map between proper cones. -/ noncomputable def map (f : E →L[ℝ] F) (K : proper_cone ℝ E) : proper_cone ℝ F := { to_convex_cone := convex_cone.closure (convex_cone.map (f : E →ₗ[ℝ] F) ↑K), nonempty' := ⟨ 0, subset_closure $ set_like.mem_coe.2 $ convex_cone.mem_map.2 @@ -152,18 +154,90 @@ lemma coe_dual (K : proper_cone ℝ E) : ↑(dual K) = (K : set E).inner_dual_co y ∈ dual K ↔ ∀ ⦃x⦄, x ∈ K → 0 ≤ ⟪x, y⟫_ℝ := by {rw [← mem_coe, coe_dual, mem_inner_dual_cone _ _], refl} --- TODO: add comap, adjoint +/-- The preimage of a proper cone under a continuous `ℝ`-linear map is a proper cone. -/ +noncomputable def comap (f : E →L[ℝ] F) (S : proper_cone ℝ F) : proper_cone ℝ E := +{ to_convex_cone := convex_cone.comap (f : E →ₗ[ℝ] F) S, + nonempty' := ⟨ 0, + begin + simp only [convex_cone.comap, mem_preimage, map_zero, set_like.mem_coe, mem_coe], + apply proper_cone.pointed, + end ⟩, + is_closed' := + begin + simp only [convex_cone.comap, continuous_linear_map.coe_coe], + apply is_closed.preimage f.2 S.is_closed, + end } + +@[simp] lemma coe_comap (f : E →L[ℝ] F) (S : proper_cone ℝ F) : (S.comap f : set E) = f ⁻¹' S := +rfl + +@[simp] lemma comap_id (S : convex_cone ℝ E) : S.comap linear_map.id = S := +set_like.coe_injective preimage_id + +lemma comap_comap (g : F →L[ℝ] G) (f : E →L[ℝ] F) (S : proper_cone ℝ G) : + (S.comap g).comap f = S.comap (g.comp f) := +set_like.coe_injective $ preimage_comp.symm + +@[simp] lemma mem_comap {f : E →L[ℝ] F} {S : proper_cone ℝ F} {x : E} : x ∈ S.comap f ↔ f x ∈ S := +iff.rfl end inner_product_space section complete_space variables {E : Type*} [normed_add_comm_group E] [inner_product_space ℝ E] [complete_space E] +variables {F : Type*} [normed_add_comm_group F] [inner_product_space ℝ F] [complete_space F] /-- The dual of the dual of a proper cone is itself. -/ @[simp] theorem dual_dual (K : proper_cone ℝ E) : K.dual.dual = K := proper_cone.ext' $ (K : convex_cone ℝ E).inner_dual_cone_of_inner_dual_cone_eq_self K.nonempty K.is_closed +/-- This is a relative version of +`convex_cone.hyperplane_separation_of_nonempty_of_is_closed_of_nmem`, which we recover by setting +`f` to be the identity map. This is a geometric interpretation of the Farkas' lemma +stated using proper cones. -/ +theorem hyperplane_separation (K : proper_cone ℝ E) {f : E →L[ℝ] F} {b : F} : + b ∈ K.map f ↔ ∀ y : F, (adjoint f y) ∈ K.dual → 0 ≤ ⟪y, b⟫_ℝ := iff.intro +begin + -- suppose `b ∈ K.map f` + simp only [proper_cone.mem_map, proper_cone.mem_dual, adjoint_inner_right, + convex_cone.mem_closure, mem_closure_iff_seq_limit], + + -- there is a sequence `seq : ℕ → F` in the image of `f` that converges to `b` + rintros ⟨seq, hmem, htends⟩ y hinner, + + suffices h : ∀ n, 0 ≤ ⟪y, seq n⟫_ℝ, from ge_of_tendsto' (continuous.seq_continuous + (continuous.inner (@continuous_const _ _ _ _ y) continuous_id) htends) h, + + intro n, + obtain ⟨_, h, hseq⟩ := hmem n, + simpa only [← hseq, real_inner_comm] using (hinner h), +end +begin + -- proof by contradiction + -- suppose `b ∉ K.map f` + intro h, + contrapose! h, + + -- as `b ∉ K.map f`, there is a hyperplane `y` separating `b` from `K.map f` + obtain ⟨y, hxy, hyb⟩ := convex_cone.hyperplane_separation_of_nonempty_of_is_closed_of_nmem _ + (K.map f).nonempty (K.map f).is_closed h, + + -- the rest of the proof is a straightforward algebraic manipulation + refine ⟨y, _, hyb⟩, + simp_rw [proper_cone.mem_dual, adjoint_inner_right], + intros x hxK, + apply hxy (f x), + rw [to_convex_cone_eq_coe, proper_cone.coe_map], + apply subset_closure, + rw [set_like.mem_coe, convex_cone.mem_map], + use ⟨x, hxK, rfl⟩, +end + +theorem hyperplane_separation_of_nmem (K : proper_cone ℝ E) {f : E →L[ℝ] F} {b : F} + (disj : b ∉ K.map f) : ∃ y : F, (adjoint f y) ∈ K.dual ∧ ⟪y, b⟫_ℝ < 0 := +by { contrapose! disj, rwa K.hyperplane_separation } + end complete_space end proper_cone From 9240e8be927a0955b9a82c6c85ef499ee3a626b8 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Thu, 29 Jun 2023 14:58:24 +0000 Subject: [PATCH 1227/1291] chore(*): add mathlib4 synchronization comments (#19218) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebraic_geometry.elliptic_curve.weierstrass` * `analysis.complex.upper_half_plane.basic` * `analysis.complex.upper_half_plane.functions_bounded_at_infty` * `analysis.complex.upper_half_plane.metric` * `analysis.complex.upper_half_plane.topology` * `analysis.fourier.riemann_lebesgue_lemma` * `analysis.inner_product_space.linear_pmap` * `analysis.special_functions.gamma.beta` * `category_theory.closed.ideal` * `category_theory.monad.equiv_mon` * `control.random` * `geometry.euclidean.monge_point` * `geometry.manifold.complex` * `geometry.manifold.partition_of_unity` * `logic.equiv.array` * `number_theory.legendre_symbol.jacobi_symbol` * `number_theory.modular` * `number_theory.modular_forms.jacobi_theta.basic` * `number_theory.modular_forms.slash_actions` * `number_theory.modular_forms.slash_invariant_forms` * `ring_theory.witt_vector.compare` * `ring_theory.witt_vector.discrete_valuation_ring` * `ring_theory.witt_vector.domain` * `ring_theory.witt_vector.frobenius` * `ring_theory.witt_vector.identities` * `ring_theory.witt_vector.init_tail` * `ring_theory.witt_vector.mul_coeff` * `ring_theory.witt_vector.truncated` * `ring_theory.witt_vector.verschiebung` * `set_theory.game.birthday` * `set_theory.game.impartial` * `set_theory.surreal.basic` * `testing.slim_check.gen` * `testing.slim_check.sampleable` * `testing.slim_check.testable` --- src/algebraic_geometry/elliptic_curve/weierstrass.lean | 3 +++ src/analysis/complex/upper_half_plane/basic.lean | 3 +++ .../complex/upper_half_plane/functions_bounded_at_infty.lean | 3 +++ src/analysis/complex/upper_half_plane/metric.lean | 3 +++ src/analysis/complex/upper_half_plane/topology.lean | 3 +++ src/analysis/fourier/riemann_lebesgue_lemma.lean | 3 +++ src/analysis/inner_product_space/linear_pmap.lean | 3 +++ src/analysis/special_functions/gamma/beta.lean | 3 +++ src/category_theory/closed/ideal.lean | 3 +++ src/category_theory/monad/equiv_mon.lean | 3 +++ src/control/random.lean | 3 +++ src/geometry/euclidean/monge_point.lean | 3 +++ src/geometry/manifold/complex.lean | 3 +++ src/geometry/manifold/partition_of_unity.lean | 3 +++ src/logic/equiv/array.lean | 3 +++ src/number_theory/legendre_symbol/jacobi_symbol.lean | 3 +++ src/number_theory/modular.lean | 3 +++ src/number_theory/modular_forms/jacobi_theta/basic.lean | 3 +++ src/number_theory/modular_forms/slash_actions.lean | 3 +++ src/number_theory/modular_forms/slash_invariant_forms.lean | 3 +++ src/ring_theory/witt_vector/compare.lean | 3 +++ src/ring_theory/witt_vector/discrete_valuation_ring.lean | 3 +++ src/ring_theory/witt_vector/domain.lean | 3 +++ src/ring_theory/witt_vector/frobenius.lean | 3 +++ src/ring_theory/witt_vector/identities.lean | 3 +++ src/ring_theory/witt_vector/init_tail.lean | 3 +++ src/ring_theory/witt_vector/mul_coeff.lean | 3 +++ src/ring_theory/witt_vector/truncated.lean | 3 +++ src/ring_theory/witt_vector/verschiebung.lean | 3 +++ src/set_theory/game/birthday.lean | 3 +++ src/set_theory/game/impartial.lean | 3 +++ src/set_theory/surreal/basic.lean | 3 +++ src/testing/slim_check/gen.lean | 3 +++ src/testing/slim_check/sampleable.lean | 3 +++ src/testing/slim_check/testable.lean | 3 +++ 35 files changed, 105 insertions(+) diff --git a/src/algebraic_geometry/elliptic_curve/weierstrass.lean b/src/algebraic_geometry/elliptic_curve/weierstrass.lean index bd0c781db4fb0..ea856a6ac27ac 100644 --- a/src/algebraic_geometry/elliptic_curve/weierstrass.lean +++ b/src/algebraic_geometry/elliptic_curve/weierstrass.lean @@ -11,6 +11,9 @@ import tactic.linear_combination /-! # Weierstrass equations of elliptic curves +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the structure of an elliptic curve as a nonsingular Weierstrass curve given by a Weierstrass equation, which is mathematically accurate in many cases but also good for computation. diff --git a/src/analysis/complex/upper_half_plane/basic.lean b/src/analysis/complex/upper_half_plane/basic.lean index b92328ac57922..02d2506086cf0 100644 --- a/src/analysis/complex/upper_half_plane/basic.lean +++ b/src/analysis/complex/upper_half_plane/basic.lean @@ -13,6 +13,9 @@ import tactic.linear_combination /-! # The upper half plane and its automorphisms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines `upper_half_plane` to be the upper half plane in `ℂ`. We furthermore equip it with the structure of an `GL_pos 2 ℝ` action by diff --git a/src/analysis/complex/upper_half_plane/functions_bounded_at_infty.lean b/src/analysis/complex/upper_half_plane/functions_bounded_at_infty.lean index 53538b4fa044d..26d6235065751 100644 --- a/src/analysis/complex/upper_half_plane/functions_bounded_at_infty.lean +++ b/src/analysis/complex/upper_half_plane/functions_bounded_at_infty.lean @@ -11,6 +11,9 @@ import order.filter.zero_and_bounded_at_filter /-! # Bounded at infinity +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + For complex valued functions on the upper half plane, this file defines the filter `at_im_infty` required for defining when functions are bounded at infinity and zero at infinity. Both of which are relevant for defining modular forms. diff --git a/src/analysis/complex/upper_half_plane/metric.lean b/src/analysis/complex/upper_half_plane/metric.lean index 39d5417caffa5..f240cfc3adbd2 100644 --- a/src/analysis/complex/upper_half_plane/metric.lean +++ b/src/analysis/complex/upper_half_plane/metric.lean @@ -10,6 +10,9 @@ import geometry.euclidean.inversion /-! # Metric on the upper half-plane +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define a `metric_space` structure on the `upper_half_plane`. We use hyperbolic (Poincaré) distance given by `dist z w = 2 * arsinh (dist (z : ℂ) w / (2 * real.sqrt (z.im * w.im)))` instead of the induced diff --git a/src/analysis/complex/upper_half_plane/topology.lean b/src/analysis/complex/upper_half_plane/topology.lean index 418466831de71..63eeb52de6c0e 100644 --- a/src/analysis/complex/upper_half_plane/topology.lean +++ b/src/analysis/complex/upper_half_plane/topology.lean @@ -13,6 +13,9 @@ import topology.homotopy.contractible /-! # Topology on the upper half plane +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we introduce a `topological_space` structure on the upper half plane and provide various instances. -/ diff --git a/src/analysis/fourier/riemann_lebesgue_lemma.lean b/src/analysis/fourier/riemann_lebesgue_lemma.lean index 31cd3a8acb7ce..d986b81e740a8 100644 --- a/src/analysis/fourier/riemann_lebesgue_lemma.lean +++ b/src/analysis/fourier/riemann_lebesgue_lemma.lean @@ -16,6 +16,9 @@ import topology.metric_space.emetric_paracompact /-! # The Riemann-Lebesgue Lemma +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove the Riemann-Lebesgue lemma, for functions on finite-dimensional real vector spaces `V`: if `f` is a function on `V` (valued in a complete normed space `E`), then the Fourier transform of `f`, viewed as a function on the dual space of `V`, tends to 0 along the diff --git a/src/analysis/inner_product_space/linear_pmap.lean b/src/analysis/inner_product_space/linear_pmap.lean index 8d36210252e4c..0613e2dfe34cd 100644 --- a/src/analysis/inner_product_space/linear_pmap.lean +++ b/src/analysis/inner_product_space/linear_pmap.lean @@ -12,6 +12,9 @@ import topology.algebra.module.basic # Partially defined linear operators on Hilbert spaces +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We will develop the basics of the theory of unbounded operators on Hilbert spaces. ## Main definitions diff --git a/src/analysis/special_functions/gamma/beta.lean b/src/analysis/special_functions/gamma/beta.lean index 42a05012901ec..9a115267c346a 100644 --- a/src/analysis/special_functions/gamma/beta.lean +++ b/src/analysis/special_functions/gamma/beta.lean @@ -11,6 +11,9 @@ import analysis.analytic.isolated_zeros /-! # The Beta function, and further properties of the Gamma function +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define the Beta integral, relate Beta and Gamma functions, and prove some refined properties of the Gamma function using these relations. diff --git a/src/category_theory/closed/ideal.lean b/src/category_theory/closed/ideal.lean index 84959823ad8ec..47f10852c5734 100644 --- a/src/category_theory/closed/ideal.lean +++ b/src/category_theory/closed/ideal.lean @@ -14,6 +14,9 @@ import category_theory.subterminal /-! # Exponential ideals +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + An exponential ideal of a cartesian closed category `C` is a subcategory `D ⊆ C` such that for any `B : D` and `A : C`, the exponential `A ⟹ B` is in `D`: resembling ring theoretic ideals. We define the notion here for inclusion functors `i : D ⥤ C` rather than explicit subcategories to diff --git a/src/category_theory/monad/equiv_mon.lean b/src/category_theory/monad/equiv_mon.lean index 4c35b16c79667..7765dc2798ed7 100644 --- a/src/category_theory/monad/equiv_mon.lean +++ b/src/category_theory/monad/equiv_mon.lean @@ -11,6 +11,9 @@ import category_theory.monoidal.Mon_ # The equivalence between `Monad C` and `Mon_ (C ⥤ C)`. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A monad "is just" a monoid in the category of endofunctors. # Definitions/Theorems diff --git a/src/control/random.lean b/src/control/random.lean index 0b6581b91df80..b6fffc4515b7d 100644 --- a/src/control/random.lean +++ b/src/control/random.lean @@ -13,6 +13,9 @@ import tactic.norm_num /-! # Rand Monad and Random Class +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This module provides tools for formulating computations guided by randomness and for defining objects that can be created randomly. diff --git a/src/geometry/euclidean/monge_point.lean b/src/geometry/euclidean/monge_point.lean index f6e741a38b287..250e4f686b38e 100644 --- a/src/geometry/euclidean/monge_point.lean +++ b/src/geometry/euclidean/monge_point.lean @@ -8,6 +8,9 @@ import geometry.euclidean.circumcenter /-! # Monge point and orthocenter +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the orthocenter of a triangle, via its n-dimensional generalization, the Monge point of a simplex. diff --git a/src/geometry/manifold/complex.lean b/src/geometry/manifold/complex.lean index 5e3d546fa2a71..bcd39594ce8bc 100644 --- a/src/geometry/manifold/complex.lean +++ b/src/geometry/manifold/complex.lean @@ -10,6 +10,9 @@ import topology.locally_constant.basic /-! # Holomorphic functions on complex manifolds +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Thanks to the rigidity of complex-differentiability compared to real-differentiability, there are many results about complex manifolds with no analogue for manifolds over a general normed field. For now, this file contains just two (closely related) such results: diff --git a/src/geometry/manifold/partition_of_unity.lean b/src/geometry/manifold/partition_of_unity.lean index 7598e00e1f479..e345b485faf25 100644 --- a/src/geometry/manifold/partition_of_unity.lean +++ b/src/geometry/manifold/partition_of_unity.lean @@ -11,6 +11,9 @@ import topology.shrinking_lemma /-! # Smooth partition of unity +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define two structures, `smooth_bump_covering` and `smooth_partition_of_unity`. Both structures describe coverings of a set by a locally finite family of supports of smooth functions with some additional properties. The former structure is mostly useful as an intermediate step in diff --git a/src/logic/equiv/array.lean b/src/logic/equiv/array.lean index 5e1f98c7d71cf..cd709ecf95cbc 100644 --- a/src/logic/equiv/array.lean +++ b/src/logic/equiv/array.lean @@ -10,6 +10,9 @@ import control.traversable.equiv /-! # Equivalences involving `array` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We keep this separate from the file containing `list`-like equivalences as those have no future in mathlib4. -/ diff --git a/src/number_theory/legendre_symbol/jacobi_symbol.lean b/src/number_theory/legendre_symbol/jacobi_symbol.lean index 597367d2409f4..7cead399fc0db 100644 --- a/src/number_theory/legendre_symbol/jacobi_symbol.lean +++ b/src/number_theory/legendre_symbol/jacobi_symbol.lean @@ -8,6 +8,9 @@ import number_theory.legendre_symbol.quadratic_reciprocity /-! # The Jacobi Symbol +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define the Jacobi symbol and prove its main properties. ## Main definitions diff --git a/src/number_theory/modular.lean b/src/number_theory/modular.lean index 901c51743781c..01ee6d1c547db 100644 --- a/src/number_theory/modular.lean +++ b/src/number_theory/modular.lean @@ -12,6 +12,9 @@ import linear_algebra.matrix.general_linear_group /-! # The action of the modular group SL(2, ℤ) on the upper half-plane +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define the action of `SL(2,ℤ)` on `ℍ` (via restriction of the `SL(2,ℝ)` action in `analysis.complex.upper_half_plane`). We then define the standard fundamental domain (`modular_group.fd`, `𝒟`) for this action and show diff --git a/src/number_theory/modular_forms/jacobi_theta/basic.lean b/src/number_theory/modular_forms/jacobi_theta/basic.lean index 5dfc5e66f6a90..86c8bd4862385 100644 --- a/src/number_theory/modular_forms/jacobi_theta/basic.lean +++ b/src/number_theory/modular_forms/jacobi_theta/basic.lean @@ -10,6 +10,9 @@ import analysis.complex.upper_half_plane.topology /-! # Jacobi's theta function +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the Jacobi theta function $$\theta(\tau) = \sum_{n \in \mathbb{Z}} \exp (i \pi n ^ 2 \tau),$$ diff --git a/src/number_theory/modular_forms/slash_actions.lean b/src/number_theory/modular_forms/slash_actions.lean index 9c980afa75e78..591978ad458c7 100644 --- a/src/number_theory/modular_forms/slash_actions.lean +++ b/src/number_theory/modular_forms/slash_actions.lean @@ -9,6 +9,9 @@ import linear_algebra.matrix.special_linear_group /-! # Slash actions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines a class of slash actions, which are families of right actions of a given group parametrized by some Type. This is modeled on the slash action of `GL_pos (fin 2) ℝ` on the space of modular forms. diff --git a/src/number_theory/modular_forms/slash_invariant_forms.lean b/src/number_theory/modular_forms/slash_invariant_forms.lean index 98b6074ce71ee..701fe0e66479a 100644 --- a/src/number_theory/modular_forms/slash_invariant_forms.lean +++ b/src/number_theory/modular_forms/slash_invariant_forms.lean @@ -8,6 +8,9 @@ import number_theory.modular_forms.slash_actions /-! # Slash invariant forms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines functions that are invariant under a `slash_action` which forms the basis for defining `modular_form` and `cusp_form`. We prove several instances for such spaces, in particular that they form a module. diff --git a/src/ring_theory/witt_vector/compare.lean b/src/ring_theory/witt_vector/compare.lean index e6e4ad3ef3fca..8041ad4f5aecb 100644 --- a/src/ring_theory/witt_vector/compare.lean +++ b/src/ring_theory/witt_vector/compare.lean @@ -12,6 +12,9 @@ import number_theory.padics.ring_homs # Comparison isomorphism between `witt_vector p (zmod p)` and `ℤ_[p]` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We construct a ring isomorphism between `witt_vector p (zmod p)` and `ℤ_[p]`. This isomorphism follows from the fact that both satisfy the universal property of the inverse limit of `zmod (p^n)`. diff --git a/src/ring_theory/witt_vector/discrete_valuation_ring.lean b/src/ring_theory/witt_vector/discrete_valuation_ring.lean index abe97d3cecbc2..f44bfbc19e7a6 100644 --- a/src/ring_theory/witt_vector/discrete_valuation_ring.lean +++ b/src/ring_theory/witt_vector/discrete_valuation_ring.lean @@ -13,6 +13,9 @@ import tactic.linear_combination # Witt vectors over a perfect ring +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file establishes that Witt vectors over a perfect field are a discrete valuation ring. When `k` is a perfect ring, a nonzero `a : 𝕎 k` can be written as `p^m * b` for some `m : ℕ` and `b : 𝕎 k` with nonzero 0th coefficient. diff --git a/src/ring_theory/witt_vector/domain.lean b/src/ring_theory/witt_vector/domain.lean index 356d1f0532bd6..2e00ed1545195 100644 --- a/src/ring_theory/witt_vector/domain.lean +++ b/src/ring_theory/witt_vector/domain.lean @@ -10,6 +10,9 @@ import ring_theory.witt_vector.identities # Witt vectors over a domain +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file builds to the proof `witt_vector.is_domain`, an instance that says if `R` is an integral domain, then so is `𝕎 R`. It depends on the API around iterated applications diff --git a/src/ring_theory/witt_vector/frobenius.lean b/src/ring_theory/witt_vector/frobenius.lean index 912a4b5962c9a..2f9692e4afaa4 100644 --- a/src/ring_theory/witt_vector/frobenius.lean +++ b/src/ring_theory/witt_vector/frobenius.lean @@ -14,6 +14,9 @@ import field_theory.perfect_closure /-! ## The Frobenius operator +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + If `R` has characteristic `p`, then there is a ring endomorphism `frobenius R p` that raises `r : R` to the power `p`. By applying `witt_vector.map` to `frobenius R p`, we obtain a ring endomorphism `𝕎 R →+* 𝕎 R`. diff --git a/src/ring_theory/witt_vector/identities.lean b/src/ring_theory/witt_vector/identities.lean index 5f30cbae076ab..e64810ad66372 100644 --- a/src/ring_theory/witt_vector/identities.lean +++ b/src/ring_theory/witt_vector/identities.lean @@ -11,6 +11,9 @@ import ring_theory.witt_vector.mul_p /-! ## Identities between operations on the ring of Witt vectors +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we derive common identities between the Frobenius and Verschiebung operators. ## Main declarations diff --git a/src/ring_theory/witt_vector/init_tail.lean b/src/ring_theory/witt_vector/init_tail.lean index 1333a32ac7a30..2be11f7e6b5d1 100644 --- a/src/ring_theory/witt_vector/init_tail.lean +++ b/src/ring_theory/witt_vector/init_tail.lean @@ -11,6 +11,9 @@ import ring_theory.witt_vector.is_poly # `init` and `tail` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given a Witt vector `x`, we are sometimes interested in its components before and after an index `n`. This file defines those operations, proves that `init` is polynomial, diff --git a/src/ring_theory/witt_vector/mul_coeff.lean b/src/ring_theory/witt_vector/mul_coeff.lean index 9903f792a5714..c5973058f8a56 100644 --- a/src/ring_theory/witt_vector/mul_coeff.lean +++ b/src/ring_theory/witt_vector/mul_coeff.lean @@ -10,6 +10,9 @@ import data.mv_polynomial.supported /-! # Leading terms of Witt vector multiplication +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The goal of this file is to study the leading terms of the formula for the `n+1`st coefficient of a product of Witt vectors `x` and `y` over a ring of characteristic `p`. We aim to isolate the `n+1`st coefficients of `x` and `y`, and express the rest of the product diff --git a/src/ring_theory/witt_vector/truncated.lean b/src/ring_theory/witt_vector/truncated.lean index 6bc9fee887a85..28a88b8f8070b 100644 --- a/src/ring_theory/witt_vector/truncated.lean +++ b/src/ring_theory/witt_vector/truncated.lean @@ -10,6 +10,9 @@ import ring_theory.witt_vector.init_tail # Truncated Witt vectors +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The ring of truncated Witt vectors (of length `n`) is a quotient of the ring of Witt vectors. It retains the first `n` coefficients of each Witt vector. In this file, we set up the basic quotient API for this ring. diff --git a/src/ring_theory/witt_vector/verschiebung.lean b/src/ring_theory/witt_vector/verschiebung.lean index f280e110e62cd..24286e1adef51 100644 --- a/src/ring_theory/witt_vector/verschiebung.lean +++ b/src/ring_theory/witt_vector/verschiebung.lean @@ -11,6 +11,9 @@ import ring_theory.witt_vector.is_poly /-! ## The Verschiebung operator +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## References * [Hazewinkel, *Witt Vectors*][Haze09] diff --git a/src/set_theory/game/birthday.lean b/src/set_theory/game/birthday.lean index c9b2642f8e8b8..163336f5f2243 100644 --- a/src/set_theory/game/birthday.lean +++ b/src/set_theory/game/birthday.lean @@ -10,6 +10,9 @@ import set_theory.ordinal.natural_ops /-! # Birthdays of games +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The birthday of a game is an ordinal that represents at which "step" the game was constructed. We define it recursively as the least ordinal larger than the birthdays of its left and right games. We prove the basic properties about these. diff --git a/src/set_theory/game/impartial.lean b/src/set_theory/game/impartial.lean index 94076d236cf67..c361208cf43b6 100644 --- a/src/set_theory/game/impartial.lean +++ b/src/set_theory/game/impartial.lean @@ -10,6 +10,9 @@ import tactic.nth_rewrite /-! # Basic definitions about impartial (pre-)games +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We will define an impartial game, one in which left and right can make exactly the same moves. Our definition differs slightly by saying that the game is always equivalent to its negative, no matter what moves are played. This allows for games such as poker-nim to be classifed as diff --git a/src/set_theory/surreal/basic.lean b/src/set_theory/surreal/basic.lean index 13d1fb908e27b..6a9d8bf27327a 100644 --- a/src/set_theory/surreal/basic.lean +++ b/src/set_theory/surreal/basic.lean @@ -10,6 +10,9 @@ import set_theory.game.ordinal /-! # Surreal numbers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The basic theory of surreal numbers, built on top of the theory of combinatorial (pre-)games. A pregame is `numeric` if all the Left options are strictly smaller than all the Right options, and diff --git a/src/testing/slim_check/gen.lean b/src/testing/slim_check/gen.lean index eff80477cdce2..d0e76abcadf90 100644 --- a/src/testing/slim_check/gen.lean +++ b/src/testing/slim_check/gen.lean @@ -11,6 +11,9 @@ import data.list.perm /-! # `gen` Monad +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This monad is used to formulate randomized computations with a parameter to specify the desired size of the result. diff --git a/src/testing/slim_check/sampleable.lean b/src/testing/slim_check/sampleable.lean index c14c317264fde..ac9b0ac097fe2 100644 --- a/src/testing/slim_check/sampleable.lean +++ b/src/testing/slim_check/sampleable.lean @@ -14,6 +14,9 @@ import tactic.linarith /-! # `sampleable` Class +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This class permits the creation samples of a given type controlling the size of those values using the `gen` monad`. It also helps minimize examples by creating smaller versions of given values. diff --git a/src/testing/slim_check/testable.lean b/src/testing/slim_check/testable.lean index 9888738ef57d3..b355cc40c8c98 100644 --- a/src/testing/slim_check/testable.lean +++ b/src/testing/slim_check/testable.lean @@ -9,6 +9,9 @@ import testing.slim_check.sampleable /-! # `testable` Class +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Testable propositions have a procedure that can generate counter-examples together with a proof that they invalidate the proposition. From 311ef8c4b4ae2804ea76b8a611bc5ea1d9c16872 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Fri, 30 Jun 2023 04:04:12 +0000 Subject: [PATCH 1228/1291] chore(category_theory/concrete_category): reorder universes (#19222) These will make life slightly easier dealing with our universe problems. Most of these changes have in fact already been made in mathlib4, and the remainder are in https://github.com/leanprover-community/mathlib4/pull/5605. Co-authored-by: Scott Morrison --- .../concrete_category/basic.lean | 54 ++++++++++--------- 1 file changed, 29 insertions(+), 25 deletions(-) diff --git a/src/category_theory/concrete_category/basic.lean b/src/category_theory/concrete_category/basic.lean index 2511901ec5bad..0029fae725175 100644 --- a/src/category_theory/concrete_category/basic.lean +++ b/src/category_theory/concrete_category/basic.lean @@ -38,7 +38,7 @@ See [Ahrens and Lumsdaine, *Displayed Categories*][ahrens2017] for related work. -/ -universes w v v' u +universes w v v' u u' namespace category_theory @@ -60,7 +60,7 @@ class concrete_category (C : Type u) [category.{v} C] := attribute [instance] concrete_category.forget_faithful /-- The forgetful functor from a concrete category to `Type u`. -/ -@[reducible] def forget (C : Type v) [category C] [concrete_category.{u} C] : C ⥤ Type u := +@[reducible] def forget (C : Type u) [category.{v} C] [concrete_category.{w} C] : C ⥤ Type w := concrete_category.forget C instance concrete_category.types : concrete_category (Type u) := @@ -75,14 +75,14 @@ You can use it on particular examples as: instance : has_coe_to_sort X := concrete_category.has_coe_to_sort X ``` -/ -def concrete_category.has_coe_to_sort (C : Type v) [category C] [concrete_category C] : - has_coe_to_sort C (Type u) := +def concrete_category.has_coe_to_sort (C : Type u) [category.{v} C] [concrete_category.{w} C] : + has_coe_to_sort C (Type w) := ⟨(concrete_category.forget C).obj⟩ section local attribute [instance] concrete_category.has_coe_to_sort -variables {C : Type v} [category C] [concrete_category C] +variables {C : Type u} [category.{v} C] [concrete_category.{w} C] @[simp] lemma forget_obj_eq_coe {X : C} : (forget C).obj X = X := rfl @@ -170,51 +170,54 @@ end `has_forget₂ C D`, where `C` and `D` are both concrete categories, provides a functor `forget₂ C D : C ⥤ D` and a proof that `forget₂ ⋙ (forget D) = forget C`. -/ -class has_forget₂ (C : Type v) (D : Type v') [category C] [concrete_category.{u} C] [category D] - [concrete_category.{u} D] := +class has_forget₂ (C : Type u) (D : Type u') [category.{v} C] [concrete_category.{w} C] + [category.{v'} D] [concrete_category.{w} D] := (forget₂ : C ⥤ D) (forget_comp : forget₂ ⋙ (forget D) = forget C . obviously) /-- The forgetful functor `C ⥤ D` between concrete categories for which we have an instance `has_forget₂ C `. -/ -@[reducible] def forget₂ (C : Type v) (D : Type v') [category C] [concrete_category C] [category D] - [concrete_category D] [has_forget₂ C D] : C ⥤ D := +@[reducible] def forget₂ (C : Type u) (D : Type u') [category.{v} C] [concrete_category.{w} C] + [category.{v'} D] [concrete_category.{w} D] [has_forget₂ C D] : C ⥤ D := has_forget₂.forget₂ -instance forget₂_faithful (C : Type v) (D : Type v') [category C] [concrete_category C] [category D] - [concrete_category D] [has_forget₂ C D] : faithful (forget₂ C D) := +instance forget₂_faithful (C : Type u) (D : Type u') [category.{v} C] [concrete_category.{w} C] + [category.{v'} D] [concrete_category.{w} D] [has_forget₂ C D] : faithful (forget₂ C D) := has_forget₂.forget_comp.faithful_of_comp -instance forget₂_preserves_monomorphisms (C : Type v) (D : Type v') [category C] - [concrete_category C] [category D] [concrete_category D] [has_forget₂ C D] +instance forget₂_preserves_monomorphisms (C : Type u) (D : Type u') + [category.{v} C] [concrete_category.{w} C] + [category.{v'} D] [concrete_category.{w} D] [has_forget₂ C D] [(forget C).preserves_monomorphisms] : (forget₂ C D).preserves_monomorphisms := have (forget₂ C D ⋙ forget D).preserves_monomorphisms, by { simp only [has_forget₂.forget_comp], apply_instance }, by exactI functor.preserves_monomorphisms_of_preserves_of_reflects _ (forget D) -instance forget₂_preserves_epimorphisms (C : Type v) (D : Type v') [category C] - [concrete_category C] [category D] [concrete_category D] [has_forget₂ C D] +instance forget₂_preserves_epimorphisms (C : Type u) (D : Type u') + [category.{v} C] [concrete_category.{w} C] + [category.{v'} D] [concrete_category.{w} D] [has_forget₂ C D] [(forget C).preserves_epimorphisms] : (forget₂ C D).preserves_epimorphisms := have (forget₂ C D ⋙ forget D).preserves_epimorphisms, by { simp only [has_forget₂.forget_comp], apply_instance }, by exactI functor.preserves_epimorphisms_of_preserves_of_reflects _ (forget D) -instance induced_category.concrete_category {C : Type v} {D : Type v'} [category D] - [concrete_category D] (f : C → D) : - concrete_category (induced_category D f) := +instance induced_category.concrete_category {C : Type u} {D : Type u'} [category.{v'} D] + [concrete_category.{w} D] (f : C → D) : + concrete_category.{w} (induced_category D f) := { forget := induced_functor f ⋙ forget D } -instance induced_category.has_forget₂ {C : Type v} {D : Type v'} [category D] [concrete_category D] +instance induced_category.has_forget₂ + {C : Type u} {D : Type u'} [category.{v'} D] [concrete_category.{w} D] (f : C → D) : has_forget₂ (induced_category D f) D := { forget₂ := induced_functor f, forget_comp := rfl } -instance full_subcategory.concrete_category {C : Type v} [category C] [concrete_category C] +instance full_subcategory.concrete_category {C : Type u} [category.{v} C] [concrete_category.{w} C] (Z : C → Prop) : concrete_category (full_subcategory Z) := { forget := full_subcategory_inclusion Z ⋙ forget C } -instance full_subcategory.has_forget₂ {C : Type v} [category C] [concrete_category C] +instance full_subcategory.has_forget₂ {C : Type u} [category.{v} C] [concrete_category.{w} C] (Z : C → Prop) : has_forget₂ (full_subcategory Z) C := { forget₂ := full_subcategory_inclusion Z, forget_comp := rfl } @@ -223,8 +226,9 @@ instance full_subcategory.has_forget₂ {C : Type v} [category C] [concrete_cate 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`. -/ -def has_forget₂.mk' {C : Type v} {D : Type v'} [category C] [concrete_category C] [category D] - [concrete_category D] (obj : C → D) (h_obj : ∀ X, (forget D).obj (obj X) = (forget C).obj X) +def has_forget₂.mk' {C : Type u} {D : Type u'} [category.{v} C] [concrete_category.{w} C] + [category.{v'} D] [concrete_category.{w} D] (obj : C → D) + (h_obj : ∀ X, (forget D).obj (obj X) = (forget C).obj X) (map : Π {X Y}, (X ⟶ Y) → (obj X ⟶ obj Y)) (h_map : ∀ {X Y} {f : X ⟶ Y}, (forget D).map (map f) == (forget C).map f) : has_forget₂ C D := @@ -233,8 +237,8 @@ has_forget₂ C D := /-- Every forgetful functor factors through the identity functor. This is not a global instance as it is prone to creating type class resolution loops. -/ -def has_forget_to_Type (C : Type v) [category C] [concrete_category C] : - has_forget₂ C (Type u) := +def has_forget_to_Type (C : Type u) [category.{v} C] [concrete_category.{w} C] : + has_forget₂ C (Type w) := { forget₂ := forget C, forget_comp := functor.comp_id _ } From 48fb5b5280e7c81672afc9524185ae994553ebf4 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Fri, 30 Jun 2023 11:29:36 +0000 Subject: [PATCH 1229/1291] refactor(*): move all `mk_simp_attribute` commands to 1 file (#19223) --- src/algebra/group/defs.lean | 4 - src/control/basic.lean | 5 +- src/control/monad/basic.lean | 4 - src/data/is_R_or_C/basic.lean | 2 - src/data/nat/parity.lean | 2 - src/data/prod/basic.lean | 4 +- src/data/set/basic.lean | 11 ++- src/data/set/image.lean | 16 +-- src/data/set/prod.lean | 10 +- src/data/subtype.lean | 2 +- src/data/typevec.lean | 5 - src/logic/basic.lean | 15 +-- src/logic/equiv/defs.lean | 17 ++-- src/logic/equiv/local_equiv.lean | 23 ----- src/logic/nontrivial.lean | 5 +- src/measure_theory/integral/bochner.lean | 11 ++- src/ring_theory/witt_vector/is_poly.lean | 3 - src/tactic/core.lean | 32 ------ src/tactic/equiv_rw.lean | 5 - src/tactic/mk_simp_attribute.lean | 119 +++++++++++++++++++++++ src/tactic/norm_cast.lean | 3 - src/tactic/split_ifs.lean | 5 - src/tactic/transport.lean | 17 ---- test/equiv_rw.lean | 6 +- 24 files changed, 173 insertions(+), 153 deletions(-) create mode 100644 src/tactic/mk_simp_attribute.lean diff --git a/src/algebra/group/defs.lean b/src/algebra/group/defs.lean index 090c13ad34f05..a33049364eb5d 100644 --- a/src/algebra/group/defs.lean +++ b/src/algebra/group/defs.lean @@ -87,10 +87,6 @@ variables {G : Type*} to the additive one. -/ -mk_simp_attribute field_simps "The simpset `field_simps` is used by the tactic `field_simp` to -reduce an expression in a field to an expression of the form `n / d` where `n` and `d` are -division-free." - section has_mul variables [has_mul G] diff --git a/src/control/basic.lean b/src/control/basic.lean index ebb5951a04034..d5ee38648ff4e 100644 --- a/src/control/basic.lean +++ b/src/control/basic.lean @@ -5,6 +5,7 @@ Authors: Johannes Hölzl Extends the theory on functors, applicatives and monads. -/ +import tactic.mk_simp_attribute universes u v w variables {α β γ : Type u} @@ -14,9 +15,6 @@ notation a ` $< `:1 f:1 := f a section functor variables {f : Type u → Type v} [functor f] [is_lawful_functor f] -run_cmd mk_simp_attr `functor_norm -run_cmd tactic.add_doc_string `simp_attr.functor_norm "Simp set for functor_norm" - @[functor_norm] theorem functor.map_map (m : α → β) (g : β → γ) (x : f α) : g <$> (m <$> x) = (g ∘ m) <$> x := (comp_map _ _ _).symm @@ -85,6 +83,7 @@ lemma seq_bind_eq (x : m α) {g : β → m γ} {f : α → β} : (f <$> x) >>= g show bind (f <$> x) g = bind x (g ∘ f), by rw [← bind_pure_comp_eq_map, bind_assoc]; simp [pure_bind] +@[monad_norm] lemma seq_eq_bind_map {x : m α} {f : m (α → β)} : f <*> x = (f >>= (<$> x)) := (bind_map_eq_seq f x).symm diff --git a/src/control/monad/basic.lean b/src/control/monad/basic.lean index 916599b5d74d1..33cfffa056f56 100644 --- a/src/control/monad/basic.lean +++ b/src/control/monad/basic.lean @@ -38,11 +38,7 @@ functor, applicative, monad, simp -/ -mk_simp_attribute monad_norm none with functor_norm - attribute [ext] reader_t.ext state_t.ext except_t.ext option_t.ext -attribute [functor_norm] bind_assoc pure_bind bind_pure -attribute [monad_norm] seq_eq_bind_map universes u v @[monad_norm] diff --git a/src/data/is_R_or_C/basic.lean b/src/data/is_R_or_C/basic.lean index c2db17c4e3424..c747fef34bf60 100644 --- a/src/data/is_R_or_C/basic.lean +++ b/src/data/is_R_or_C/basic.lean @@ -70,8 +70,6 @@ class is_R_or_C (K : Type*) end -mk_simp_attribute is_R_or_C_simps "Simp attribute for lemmas about `is_R_or_C`" - variables {K E : Type*} [is_R_or_C K] namespace is_R_or_C diff --git a/src/data/nat/parity.lean b/src/data/nat/parity.lean index 6a0cf0a24bcbc..5ec996b4dfb0b 100644 --- a/src/data/nat/parity.lean +++ b/src/data/nat/parity.lean @@ -91,8 +91,6 @@ mod_two_add_add_odd_mod_two m odd_one @[simp] theorem succ_mod_two_add_mod_two (m : ℕ) : (m + 1) % 2 + m % 2 = 1 := by rw [add_comm, mod_two_add_succ_mod_two] -mk_simp_attribute parity_simps "Simp attribute for lemmas about `even`" - @[simp] theorem not_even_one : ¬ even 1 := by rw even_iff; norm_num diff --git a/src/data/prod/basic.lean b/src/data/prod/basic.lean index 240beda2eb6e2..a2a159aced62e 100644 --- a/src/data/prod/basic.lean +++ b/src/data/prod/basic.lean @@ -37,7 +37,9 @@ prod.exists @[simp] lemma fst_comp_mk (x : α) : prod.fst ∘ (prod.mk x : β → α × β) = function.const β x := rfl -@[simp] lemma map_mk (f : α → γ) (g : β → δ) (a : α) (b : β) : map f g (a, b) = (f a, g b) := rfl +@[simp, mfld_simps] lemma map_mk (f : α → γ) (g : β → δ) (a : α) (b : β) : + map f g (a, b) = (f a, g b) := +rfl lemma map_fst (f : α → γ) (g : β → δ) (p : α × β) : (map f g p).1 = f (p.1) := rfl diff --git a/src/data/set/basic.lean b/src/data/set/basic.lean index 08af3259a3f2e..9ac1590fab649 100644 --- a/src/data/set/basic.lean +++ b/src/data/set/basic.lean @@ -409,7 +409,7 @@ Mathematically it is the same as `α` but it has a different type. @[simp] theorem set_of_true : {x : α | true} = univ := rfl -@[simp] theorem mem_univ (x : α) : x ∈ @univ α := trivial +@[simp, mfld_simps] theorem mem_univ (x : α) : x ∈ @univ α := trivial @[simp] lemma univ_eq_empty_iff : (univ : set α) = ∅ ↔ is_empty α := eq_empty_iff_forall_not_mem.trans ⟨λ H, ⟨λ x, H x trivial⟩, λ H x _, @is_empty.false α H x⟩ @@ -541,7 +541,8 @@ by simp only [← subset_empty_iff]; exact union_subset_iff theorem inter_def {s₁ s₂ : set α} : s₁ ∩ s₂ = {a | a ∈ s₁ ∧ a ∈ s₂} := rfl -@[simp] theorem mem_inter_iff (x : α) (a b : set α) : x ∈ a ∩ b ↔ (x ∈ a ∧ x ∈ b) := iff.rfl +@[simp, mfld_simps] +theorem mem_inter_iff (x : α) (a b : set α) : x ∈ a ∩ b ↔ (x ∈ a ∧ x ∈ b) := iff.rfl theorem mem_inter {x : α} {a b : set α} (ha : x ∈ a) (hb : x ∈ b) : x ∈ a ∩ b := ⟨ha, hb⟩ @@ -569,7 +570,7 @@ ext $ λ x, and.left_comm theorem inter_right_comm (s₁ s₂ s₃ : set α) : (s₁ ∩ s₂) ∩ s₃ = (s₁ ∩ s₃) ∩ s₂ := ext $ λ x, and.right_comm -@[simp] theorem inter_subset_left (s t : set α) : s ∩ t ⊆ s := λ x, and.left +@[simp, mfld_simps] theorem inter_subset_left (s t : set α) : s ∩ t ⊆ s := λ x, and.left @[simp] theorem inter_subset_right (s t : set α) : s ∩ t ⊆ t := λ x, and.right @@ -596,9 +597,9 @@ lemma inter_congr_right (hs : t ∩ u ⊆ s) (ht : s ∩ u ⊆ t) : s ∩ u = t lemma inter_eq_inter_iff_left : s ∩ t = s ∩ u ↔ s ∩ u ⊆ t ∧ s ∩ t ⊆ u := inf_eq_inf_iff_left lemma inter_eq_inter_iff_right : s ∩ u = t ∩ u ↔ t ∩ u ⊆ s ∧ s ∩ u ⊆ t := inf_eq_inf_iff_right -@[simp] theorem inter_univ (a : set α) : a ∩ univ = a := inf_top_eq +@[simp, mfld_simps] theorem inter_univ (a : set α) : a ∩ univ = a := inf_top_eq -@[simp] theorem univ_inter (a : set α) : univ ∩ a = a := top_inf_eq +@[simp, mfld_simps] theorem univ_inter (a : set α) : univ ∩ a = a := top_inf_eq theorem inter_subset_inter {s₁ s₂ t₁ t₂ : set α} (h₁ : s₁ ⊆ t₁) (h₂ : s₂ ⊆ t₂) : s₁ ∩ s₂ ⊆ t₁ ∩ t₂ := λ x, and.imp (@h₁ _) (@h₂ _) diff --git a/src/data/set/image.lean b/src/data/set/image.lean index fba4f00b8d959..4e942d43a7266 100644 --- a/src/data/set/image.lean +++ b/src/data/set/image.lean @@ -50,7 +50,7 @@ variables {f : α → β} {g : β → γ} @[simp] theorem preimage_empty : f ⁻¹' ∅ = ∅ := rfl -@[simp] theorem mem_preimage {s : set β} {a : α} : (a ∈ f ⁻¹' s) ↔ (f a ∈ s) := iff.rfl +@[simp, mfld_simps] theorem mem_preimage {s : set β} {a : α} : (a ∈ f ⁻¹' s) ↔ (f a ∈ s) := iff.rfl lemma preimage_congr {f g : α → β} {s : set β} (h : ∀ (x : α), f x = g x) : f ⁻¹' s = g ⁻¹' s := by { congr' with x, apply_assumption } @@ -58,11 +58,11 @@ by { congr' with x, apply_assumption } theorem preimage_mono {s t : set β} (h : s ⊆ t) : f ⁻¹' s ⊆ f ⁻¹' t := assume x hx, h hx -@[simp] theorem preimage_univ : f ⁻¹' univ = univ := rfl +@[simp, mfld_simps] theorem preimage_univ : f ⁻¹' univ = univ := rfl theorem subset_preimage_univ {s : set α} : s ⊆ f ⁻¹' univ := subset_univ _ -@[simp] theorem preimage_inter {s t : set β} : f ⁻¹' (s ∩ t) = f ⁻¹' s ∩ f ⁻¹' t := rfl +@[simp, mfld_simps] theorem preimage_inter {s t : set β} : f ⁻¹' (s ∩ t) = f ⁻¹' s ∩ f ⁻¹' t := rfl @[simp] theorem preimage_union {s t : set β} : f ⁻¹' (s ∪ t) = f ⁻¹' s ∪ f ⁻¹' t := rfl @@ -80,7 +80,7 @@ rfl @[simp] lemma preimage_id_eq : preimage (id : α → α) = id := rfl -theorem preimage_id {s : set α} : id ⁻¹' s = s := rfl +@[mfld_simps] theorem preimage_id {s : set α} : id ⁻¹' s = s := rfl @[simp] theorem preimage_id' {s : set α} : (λ x, x) ⁻¹' s = s := rfl @@ -152,6 +152,7 @@ theorem mem_image_iff_bex {f : α → β} {s : set α} {y : β} : lemma image_eta (f : α → β) : f '' s = (λ x, f x) '' s := rfl +@[mfld_simps] theorem mem_image_of_mem (f : α → β) {x : α} {a : set α} (h : x ∈ a) : f x ∈ f '' a := ⟨_, h, rfl⟩ @@ -252,7 +253,8 @@ by { ext, simp [image, eq_comm] } ext $ λ x, ⟨λ ⟨y, _, h⟩, h ▸ mem_singleton _, λ h, (eq_of_mem_singleton h).symm ▸ hs.imp (λ y hy, ⟨hy, rfl⟩)⟩ -@[simp] lemma image_eq_empty {α β} {f : α → β} {s : set α} : f '' s = ∅ ↔ s = ∅ := +@[simp, mfld_simps] +lemma image_eq_empty {α β} {f : α → β} {s : set α} : f '' s = ∅ ↔ s = ∅ := by { simp only [eq_empty_iff_forall_not_mem], exact ⟨λ H a ha, H _ ⟨_, ha, rfl⟩, λ H b ⟨_, ha, _⟩, H _ ha⟩ } @@ -506,7 +508,7 @@ def range (f : ι → α) : set α := {x | ∃y, f y = x} @[simp] theorem mem_range {x : α} : x ∈ range f ↔ ∃ y, f y = x := iff.rfl -@[simp] theorem mem_range_self (i : ι) : f i ∈ range f := ⟨i, rfl⟩ +@[simp, mfld_simps] theorem mem_range_self (i : ι) : f i ∈ range f := ⟨i, rfl⟩ theorem forall_range_iff {p : α → Prop} : (∀ a ∈ range f, p a) ↔ (∀ i, p (f i)) := by simp @@ -649,7 +651,7 @@ theorem preimage_image_preimage {f : α → β} {s : set β} : f ⁻¹' (f '' (f ⁻¹' s)) = f ⁻¹' s := by rw [image_preimage_eq_inter_range, preimage_inter_range] -@[simp] theorem range_id : range (@id α) = univ := range_iff_surjective.2 surjective_id +@[simp, mfld_simps] theorem range_id : range (@id α) = univ := range_iff_surjective.2 surjective_id @[simp] theorem range_id' : range (λ (x : α), x) = univ := range_id diff --git a/src/data/set/prod.lean b/src/data/set/prod.lean index e6f35bb6ad86c..90f2736c294bc 100644 --- a/src/data/set/prod.lean +++ b/src/data/set/prod.lean @@ -42,9 +42,9 @@ lemma prod_eq (s : set α) (t : set β) : s ×ˢ t = prod.fst ⁻¹' s ∩ prod. lemma mem_prod_eq {p : α × β} : p ∈ s ×ˢ t = (p.1 ∈ s ∧ p.2 ∈ t) := rfl -@[simp] lemma mem_prod {p : α × β} : p ∈ s ×ˢ t ↔ p.1 ∈ s ∧ p.2 ∈ t := iff.rfl +@[simp, mfld_simps] lemma mem_prod {p : α × β} : p ∈ s ×ˢ t ↔ p.1 ∈ s ∧ p.2 ∈ t := iff.rfl -@[simp] lemma prod_mk_mem_set_prod_eq : (a, b) ∈ s ×ˢ t = (a ∈ s ∧ b ∈ t) := rfl +@[simp, mfld_simps] lemma prod_mk_mem_set_prod_eq : (a, b) ∈ s ×ˢ t = (a ∈ s ∧ b ∈ t) := rfl lemma mk_mem_prod (ha : a ∈ s) (hb : b ∈ t) : (a, b) ∈ s ×ˢ t := ⟨ha, hb⟩ @@ -77,7 +77,7 @@ by simp [and_assoc] @[simp] lemma empty_prod : (∅ : set α) ×ˢ t = ∅ := by { ext, exact false_and _ } -@[simp] lemma univ_prod_univ : @univ α ×ˢ @univ β = univ := by { ext, exact true_and _ } +@[simp, mfld_simps] lemma univ_prod_univ : @univ α ×ˢ @univ β = univ := by { ext, exact true_and _ } lemma univ_prod {t : set β} : (univ : set α) ×ˢ t = prod.snd ⁻¹' t := by simp [prod_eq] @@ -103,6 +103,7 @@ by { ext ⟨x, y⟩, simp only [←and_and_distrib_right, mem_inter_iff, mem_pro lemma prod_inter : s ×ˢ (t₁ ∩ t₂) = s ×ˢ t₁ ∩ s ×ˢ t₂ := by { ext ⟨x, y⟩, simp only [←and_and_distrib_left, mem_inter_iff, mem_prod] } +@[mfld_simps] lemma prod_inter_prod : s₁ ×ˢ t₁ ∩ s₂ ×ˢ t₂ = (s₁ ∩ s₂) ×ˢ (t₁ ∩ t₂) := by { ext ⟨x, y⟩, simp [and_assoc, and.left_comm] } @@ -186,7 +187,7 @@ lemma prod_range_range_eq {m₁ : α → γ} {m₂ : β → δ} : (range m₁) ×ˢ (range m₂) = range (λ p : α × β, (m₁ p.1, m₂ p.2)) := ext $ by simp [range] -@[simp] lemma range_prod_map {m₁ : α → γ} {m₂ : β → δ} : +@[simp, mfld_simps] lemma range_prod_map {m₁ : α → γ} {m₂ : β → δ} : range (prod.map m₁ m₂) = (range m₁) ×ˢ (range m₂) := prod_range_range_eq.symm @@ -219,7 +220,6 @@ lemma prod_sub_preimage_iff {W : set γ} {f : α × β → γ} : s ×ˢ t ⊆ f ⁻¹' W ↔ ∀ a b, a ∈ s → b ∈ t → f (a, b) ∈ W := by simp [subset_def] - lemma image_prod_mk_subset_prod {f : α → β} {g : α → γ} {s : set α} : (λ x, (f x, g x)) '' s ⊆ (f '' s) ×ˢ (g '' s) := by { rintros _ ⟨x, hx, rfl⟩, exact mk_mem_prod (mem_image_of_mem f hx) (mem_image_of_mem g hx) } diff --git a/src/data/subtype.lean b/src/data/subtype.lean index 5dd326422b4cb..10dc89a45d6f0 100644 --- a/src/data/subtype.lean +++ b/src/data/subtype.lean @@ -84,7 +84,7 @@ ext_iff @[simp] theorem coe_eta (a : {a // p a}) (h : p a) : mk ↑a h = a := subtype.ext rfl -@[simp] theorem coe_mk (a h) : (@mk α p a h : α) = a := rfl +@[simp, mfld_simps] theorem coe_mk (a h) : (@mk α p a h : α) = a := rfl @[simp, nolint simp_nf] -- built-in reduction doesn't always work theorem mk_eq_mk {a h a' h'} : @mk α p a h = @mk α p a' h' ↔ a = a' := diff --git a/src/data/typevec.lean b/src/data/typevec.lean index 60416b60ebafb..55224d6e7ef0b 100644 --- a/src/data/typevec.lean +++ b/src/data/typevec.lean @@ -235,11 +235,6 @@ eq_of_drop_last_eq rfl rfl instance subsingleton0 : subsingleton (typevec 0) := ⟨ λ a b, funext $ λ a, fin2.elim0 a ⟩ -run_cmd do - mk_simp_attr `typevec, - tactic.add_doc_string `simp_attr.typevec -"simp set for the manipulation of typevec and arrow expressions" - local prefix `♯`:0 := cast (by try { simp }; congr' 1; try { simp }) /-- cases distinction for 0-length type vector -/ diff --git a/src/logic/basic.lean b/src/logic/basic.lean index e4b5d9cb3afe4..71baa00527e9d 100644 --- a/src/logic/basic.lean +++ b/src/logic/basic.lean @@ -3,7 +3,7 @@ Copyright (c) 2016 Jeremy Avigad. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Leonardo de Moura -/ -import tactic.doc_commands +import tactic.mk_simp_attribute import tactic.reserved_notation /-! @@ -63,7 +63,7 @@ instance psum.inhabited_right {α β} [inhabited β] : inhabited (psum α β) := {α} [subsingleton α] : decidable_eq α | a b := is_true (subsingleton.elim a b) -@[simp] lemma eq_iff_true_of_subsingleton {α : Sort*} [subsingleton α] (x y : α) : +@[simp, nontriviality] lemma eq_iff_true_of_subsingleton {α : Sort*} [subsingleton α] (x y : α) : x = y ↔ true := by cc @@ -274,7 +274,7 @@ theorem imp_and_distrib {α} : (α → b ∧ c) ↔ (α → b) ∧ (α → c) := ⟨λ h, ⟨λ ha, (h ha).left, λ ha, (h ha).right⟩, λ h ha, ⟨h.left ha, h.right ha⟩⟩ -@[simp] theorem and_imp : (a ∧ b → c) ↔ (a → b → c) := +@[simp, mfld_simps] theorem and_imp : (a ∧ b → c) ↔ (a → b → c) := iff.intro (λ h ha hb, h ⟨ha, hb⟩) (λ h ⟨ha, hb⟩, h ha hb) theorem iff_def : (a ↔ b) ↔ (a → b) ∧ (b → a) := @@ -842,7 +842,7 @@ end mem section equality variables {α : Sort*} {a b : α} -@[simp] theorem heq_iff_eq : a == b ↔ a = b := +@[simp, mfld_simps] theorem heq_iff_eq : a == b ↔ a = b := ⟨eq_of_heq, heq_of_eq⟩ theorem proof_irrel_heq {p q : Prop} (hp : p) (hq : q) : hp == hq := @@ -865,12 +865,12 @@ theorem eq_equivalence : equivalence (@eq α) := ⟨eq.refl, @eq.symm _, @eq.trans _⟩ /-- Transport through trivial families is the identity. -/ -@[simp] +@[simp, transport_simps] lemma eq_rec_constant {α : Sort*} {a a' : α} {β : Sort*} (y : β) (h : a = a') : (@eq.rec α a (λ a, β) y a' h) = y := by { cases h, refl, } -@[simp] +@[simp, transport_simps] lemma eq_mp_eq_cast {α β : Sort*} (h : α = β) : eq.mp h = cast h := rfl @[simp] @@ -1096,6 +1096,7 @@ let ⟨a⟩ := ha in (λ hb, hb $ h $ λ x, (not_imp.1 (h' x)).1), λ ⟨x, hx⟩ h, hx (h x)⟩ -- TODO: duplicate of a lemma in core +@[mfld_simps] theorem forall_true_iff : (α → true) ↔ true := implies_true_iff α @@ -1118,7 +1119,7 @@ exists.elim h (λ x hx, ⟨x, and.left hx⟩) (∃! x, p x) ↔ ∃ x, p x := ⟨λ h, h.exists, Exists.imp $ λ x hx, ⟨hx, λ y _, subsingleton.elim y x⟩⟩ -@[simp] theorem forall_const (α : Sort*) [i : nonempty α] : (α → b) ↔ b := +@[simp, mfld_simps] theorem forall_const (α : Sort*) [i : nonempty α] : (α → b) ↔ b := ⟨i.elim, λ hb x, hb⟩ /-- For some reason simp doesn't use `forall_const` to simplify in this case. -/ diff --git a/src/logic/equiv/defs.lean b/src/logic/equiv/defs.lean index abb46c6753b6c..490c633c499db 100644 --- a/src/logic/equiv/defs.lean +++ b/src/logic/equiv/defs.lean @@ -120,10 +120,10 @@ initialize_simps_projections equiv (to_fun → apply, inv_fun → symm_apply) @[trans] protected def trans (e₁ : α ≃ β) (e₂ : β ≃ γ) : α ≃ γ := ⟨e₂ ∘ e₁, e₁.symm ∘ e₂.symm, e₂.left_inv.comp e₁.left_inv, e₂.right_inv.comp e₁.right_inv⟩ -@[simp] +@[simp, transport_simps, mfld_simps] lemma to_fun_as_coe (e : α ≃ β) : e.to_fun = e := rfl -@[simp] +@[simp, mfld_simps] lemma inv_fun_as_coe (e : α ≃ β) : e.inv_fun = e.symm := rfl protected theorem injective (e : α ≃ β) : injective e := equiv_like.injective e @@ -189,10 +189,11 @@ theorem refl_apply (x : α) : equiv.refl α x = x := rfl theorem trans_apply (f : α ≃ β) (g : β ≃ γ) (a : α) : (f.trans g) a = g (f a) := rfl -@[simp] theorem apply_symm_apply (e : α ≃ β) (x : β) : e (e.symm x) = x := +@[simp, equiv_rw_simp] theorem apply_symm_apply (e : α ≃ β) (x : β) : e (e.symm x) = x := e.right_inv x -@[simp] theorem symm_apply_apply (e : α ≃ β) (x : α) : e.symm (e x) = x := +@[simp, equiv_rw_simp, transport_simps] +theorem symm_apply_apply (e : α ≃ β) (x : α) : e.symm (e x) = x := e.left_inv x @[simp] theorem symm_comp_self (e : α ≃ β) : e.symm ∘ e = id := funext e.symm_apply_apply @@ -208,6 +209,7 @@ e.left_inv x theorem apply_eq_iff_eq (f : α ≃ β) {x y : α} : f x = f y ↔ x = y := equiv_like.apply_eq_iff_eq f +@[transport_simps] theorem apply_eq_iff_eq_symm_apply {α β : Sort*} (f : α ≃ β) {x : α} {y : β} : f x = y ↔ x = f.symm y := begin @@ -234,7 +236,7 @@ lemma symm_apply_eq {α β} (e : α ≃ β) {x y} : e.symm x = y ↔ x = e y := lemma eq_symm_apply {α β} (e : α ≃ β) {x y} : y = e.symm x ↔ e y = x := (eq_comm.trans e.symm_apply_eq).trans eq_comm -@[simp] theorem symm_symm (e : α ≃ β) : e.symm.symm = e := by { cases e, refl } +@[simp, equiv_rw_simp] theorem symm_symm (e : α ≃ β) : e.symm.symm = e := by { cases e, refl } @[simp] theorem trans_refl (e : α ≃ β) : e.trans (equiv.refl β) = e := by { cases e, refl } @@ -411,7 +413,7 @@ A version of `equiv.arrow_congr` in `Type`, rather than `Sort`. The `equiv_rw` tactic is not able to use the default `Sort` level `equiv.arrow_congr`, because Lean's universe rules will not unify `?l_1` with `imax (1 ?m_1)`. -/ -@[congr, simps apply] +@[congr, simps apply { attrs := [`simp, `transport_simps] }] def arrow_congr' {α₁ β₁ α₂ β₂ : Type*} (hα : α₁ ≃ α₂) (hβ : β₁ ≃ β₂) : (α₁ → β₁) ≃ (α₂ → β₂) := equiv.arrow_congr hα hβ @@ -638,7 +640,8 @@ def sigma_congr {α₁ α₂} {β₁ : α₁ → Sort*} {β₂ : α₂ → Sort* (sigma_congr_right F).trans (sigma_congr_left f) /-- `sigma` type with a constant fiber is equivalent to the product. -/ -@[simps apply symm_apply] def sigma_equiv_prod (α β : Type*) : (Σ_:α, β) ≃ α × β := +@[simps apply symm_apply { attrs := [`simp, `mfld_simps] }] +def sigma_equiv_prod (α β : Type*) : (Σ_:α, β) ≃ α × β := ⟨λ a, ⟨a.1, a.2⟩, λ a, ⟨a.1, a.2⟩, λ ⟨a, b⟩, rfl, λ ⟨a, b⟩, rfl⟩ /-- If each fiber of a `sigma` type is equivalent to a fixed type, then the sigma type diff --git a/src/logic/equiv/local_equiv.lean b/src/logic/equiv/local_equiv.lean index 7b1c693559b8f..d86e01abfeddc 100644 --- a/src/logic/equiv/local_equiv.lean +++ b/src/logic/equiv/local_equiv.lean @@ -68,29 +68,6 @@ then it should use `e.source ∩ s` or `e.target ∩ t`, not `s ∩ e.source` or -/ -mk_simp_attribute mfld_simps "The simpset `mfld_simps` records several simp lemmas that are -especially useful in manifolds. It is a subset of the whole set of simp lemmas, but it makes it -possible to have quicker proofs (when used with `squeeze_simp` or `simp only`) while retaining -readability. - -The typical use case is the following, in a file on manifolds: -If `simp [foo, bar]` is slow, replace it with `squeeze_simp [foo, bar] with mfld_simps` and paste -its output. The list of lemmas should be reasonable (contrary to the output of -`squeeze_simp [foo, bar]` which might contain tens of lemmas), and the outcome should be quick -enough. -" - --- register in the simpset `mfld_simps` several lemmas that are often useful when dealing --- with manifolds -attribute [mfld_simps] id.def function.comp.left_id set.mem_set_of_eq set.image_eq_empty -set.univ_inter set.preimage_univ set.prod_mk_mem_set_prod_eq and_true set.mem_univ -set.mem_image_of_mem true_and set.mem_inter_iff set.mem_preimage function.comp_app -set.inter_subset_left set.mem_prod set.range_id set.range_prod_map and_self set.mem_range_self -eq_self_iff_true forall_const forall_true_iff set.inter_univ set.preimage_id function.comp.right_id -not_false_iff and_imp set.prod_inter_prod set.univ_prod_univ true_or or_true prod.map_mk -set.preimage_inter heq_iff_eq equiv.sigma_equiv_prod_apply equiv.sigma_equiv_prod_symm_apply -subtype.coe_mk equiv.to_fun_as_coe equiv.inv_fun_as_coe - /-- Common `@[simps]` configuration options used for manifold-related declarations. -/ def mfld_cfg : simps_cfg := {attrs := [`simp, `mfld_simps], fully_applied := ff} diff --git a/src/logic/nontrivial.lean b/src/logic/nontrivial.lean index 99388b9a6b5a1..14ad9c3c2116d 100644 --- a/src/logic/nontrivial.lean +++ b/src/logic/nontrivial.lean @@ -178,13 +178,10 @@ end pi instance function.nontrivial [h : nonempty α] [nontrivial β] : nontrivial (α → β) := h.elim $ λ a, pi.nontrivial_at a -mk_simp_attribute nontriviality "Simp lemmas for `nontriviality` tactic" - +@[nontriviality] protected lemma subsingleton.le [preorder α] [subsingleton α] (x y : α) : x ≤ y := le_of_eq (subsingleton.elim x y) -attribute [nontriviality] eq_iff_true_of_subsingleton subsingleton.le - namespace bool instance : nontrivial bool := ⟨⟨tt,ff, tt_eq_ff_eq_false⟩⟩ diff --git a/src/measure_theory/integral/bochner.lean b/src/measure_theory/integral/bochner.lean index c594e1e066aed..84076d698e9b8 100644 --- a/src/measure_theory/integral/bochner.lean +++ b/src/measure_theory/integral/bochner.lean @@ -638,24 +638,28 @@ end variables {α E} +@[integral_simps] lemma integral_add (f g : α →₁[μ] E) : integral (f + g) = integral f + integral g := begin simp only [integral], exact map_add integral_clm f g end +@[integral_simps] lemma integral_neg (f : α →₁[μ] E) : integral (-f) = - integral f := begin simp only [integral], exact map_neg integral_clm f end +@[integral_simps] lemma integral_sub (f g : α →₁[μ] E) : integral (f - g) = integral f - integral g := begin simp only [integral], exact map_sub integral_clm f g end +@[integral_simps] lemma integral_smul (c : 𝕜) (f : α →₁[μ] E) : integral (c • f) = c • integral f := begin simp only [integral], @@ -797,6 +801,7 @@ begin exact set_to_fun_finset_sum (dominated_fin_meas_additive_weighted_smul _) s hf end +@[integral_simps] lemma integral_neg (f : α → E) : ∫ a, -f a ∂μ = - ∫ a, f a ∂μ := begin simp only [integral, L1.integral], @@ -817,6 +822,7 @@ lemma integral_sub' (hf : integrable f μ) (hg : integrable g μ) : ∫ a, (f - g) a ∂μ = ∫ a, f a ∂μ - ∫ a, g a ∂μ := integral_sub hf hg +@[integral_simps] lemma integral_smul (c : 𝕜) (f : α → E) : ∫ a, c • (f a) ∂μ = c • ∫ a, f a ∂μ := begin @@ -1664,11 +1670,6 @@ end end properties -mk_simp_attribute integral_simps "Simp set for integral rules." - -attribute [integral_simps] integral_neg integral_smul L1.integral_add L1.integral_sub - L1.integral_smul L1.integral_neg - section integral_trim variables {H β γ : Type*} [normed_add_comm_group H] diff --git a/src/ring_theory/witt_vector/is_poly.lean b/src/ring_theory/witt_vector/is_poly.lean index 5be39c921c856..f967401b8e131 100644 --- a/src/ring_theory/witt_vector/is_poly.lean +++ b/src/ring_theory/witt_vector/is_poly.lean @@ -100,9 +100,6 @@ end We define it here so it is a shared import. -/ -mk_simp_attribute ghost_simps -"Simplification rules for ghost equations" - namespace tactic namespace interactive setup_tactic_parser diff --git a/src/tactic/core.lean b/src/tactic/core.lean index 845f0254d8fc3..2364d7bd99a08 100644 --- a/src/tactic/core.lean +++ b/src/tactic/core.lean @@ -2409,38 +2409,6 @@ add_tactic_doc decl_names := [`tactic.import_private_cmd], tags := ["renaming"] } -/-- -The command `mk_simp_attribute simp_name "description"` creates a simp set with name `simp_name`. -Lemmas tagged with `@[simp_name]` will be included when `simp with simp_name` is called. -`mk_simp_attribute simp_name none` will use a default description. - -Appending the command with `with attr1 attr2 ...` will include all declarations tagged with -`attr1`, `attr2`, ... in the new simp set. - -This command is preferred to using ``run_cmd mk_simp_attr `simp_name`` since it adds a doc string -to the attribute that is defined. If you need to create a simp set in a file where this command is -not available, you should use -```lean -run_cmd mk_simp_attr `simp_name -run_cmd add_doc_string `simp_attr.simp_name "Description of the simp set here" -``` --/ -@[user_command] -meta def mk_simp_attribute_cmd (_ : parse $ tk "mk_simp_attribute") : lean.parser unit := -do n ← ident, - d ← parser.pexpr, - d ← to_expr ``(%%d : option string), - descr ← eval_expr (option string) d, - with_list ← (tk "with" *> many ident) <|> return [], - mk_simp_attr n with_list, - add_doc_string (name.append `simp_attr n) $ descr.get_or_else $ "simp set for " ++ to_string n - -add_tactic_doc -{ name := "mk_simp_attribute", - category := doc_category.cmd, - decl_names := [`tactic.mk_simp_attribute_cmd], - tags := ["simplification"] } - /-- Given a user attribute name `attr_name`, `get_user_attribute_name attr_name` returns the name of the declaration that defines this attribute. diff --git a/src/tactic/equiv_rw.lean b/src/tactic/equiv_rw.lean index 766703bf8924f..ca74919c3eecc 100644 --- a/src/tactic/equiv_rw.lean +++ b/src/tactic/equiv_rw.lean @@ -187,11 +187,6 @@ do -- to compress away some `map_equiv equiv.refl` subexpressions. prod.fst <$> new_eqv.simp {fail_if_unchanged := ff} -mk_simp_attribute equiv_rw_simp "The simpset `equiv_rw_simp` is used by the tactic `equiv_rw` to -simplify applications of equivalences and their inverses." - -attribute [equiv_rw_simp] equiv.symm_symm equiv.apply_symm_apply equiv.symm_apply_apply - /-- Attempt to replace the hypothesis with name `x` by transporting it along the equivalence in `e : α ≃ β`. diff --git a/src/tactic/mk_simp_attribute.lean b/src/tactic/mk_simp_attribute.lean new file mode 100644 index 0000000000000..2422f9675391a --- /dev/null +++ b/src/tactic/mk_simp_attribute.lean @@ -0,0 +1,119 @@ +/- +Copyright (c) 2019 Rob Lewis All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Rob Lewis +-/ +import tactic.doc_commands +/-! +# User command to register `simp` attributes + +In this file we define a command `mk_simp_attribute` that can be used to register `simp` sets. We +also define all `simp` attributes that are used in the library and tag lemmas from Lean core with +these attributes. +-/ + +/-! +### User command +-/ + +section cmd + +open interactive lean lean.parser + +namespace tactic + +/-- +The command `mk_simp_attribute simp_name "description"` creates a simp set with name `simp_name`. +Lemmas tagged with `@[simp_name]` will be included when `simp with simp_name` is called. +`mk_simp_attribute simp_name none` will use a default description. + +Appending the command with `with attr1 attr2 ...` will include all declarations tagged with +`attr1`, `attr2`, ... in the new simp set. + +This command is preferred to using ``run_cmd mk_simp_attr `simp_name`` since it adds a doc string +to the attribute that is defined. If you need to create a simp set in a file where this command is +not available, you should use +```lean +run_cmd mk_simp_attr `simp_name +run_cmd add_doc_string `simp_attr.simp_name "Description of the simp set here" +``` +-/ +@[user_command] +meta def mk_simp_attribute_cmd (_ : parse $ tk "mk_simp_attribute") : lean.parser unit := +do n ← ident, + d ← parser.pexpr, + d ← to_expr ``(%%d : option string), + descr ← eval_expr (option string) d, + with_list ← (tk "with" *> many ident) <|> return [], + mk_simp_attr n with_list, + add_doc_string (name.append `simp_attr n) $ descr.get_or_else $ "simp set for " ++ to_string n + +add_tactic_doc +{ name := "mk_simp_attribute", + category := doc_category.cmd, + decl_names := [`tactic.mk_simp_attribute_cmd], + tags := ["simplification"] } + +end tactic + +end cmd + +/-! +### Attributes +-/ + +mk_simp_attribute equiv_rw_simp "The simpset `equiv_rw_simp` is used by the tactic `equiv_rw` to +simplify applications of equivalences and their inverses." + +mk_simp_attribute field_simps "The simpset `field_simps` is used by the tactic `field_simp` to +reduce an expression in a field to an expression of the form `n / d` where `n` and `d` are +division-free." + +mk_simp_attribute functor_norm "Simp set for functor_norm" + +attribute [functor_norm] bind_assoc pure_bind bind_pure + +mk_simp_attribute ghost_simps "Simplification rules for ghost equations" + +mk_simp_attribute integral_simps "Simp set for integral rules." + +mk_simp_attribute is_R_or_C_simps "Simp attribute for lemmas about `is_R_or_C`" + +mk_simp_attribute mfld_simps "The simpset `mfld_simps` records several simp lemmas that are +especially useful in manifolds. It is a subset of the whole set of simp lemmas, but it makes it +possible to have quicker proofs (when used with `squeeze_simp` or `simp only`) while retaining +readability. + +The typical use case is the following, in a file on manifolds: +If `simp [foo, bar]` is slow, replace it with `squeeze_simp [foo, bar] with mfld_simps` and paste +its output. The list of lemmas should be reasonable (contrary to the output of +`squeeze_simp [foo, bar]` which might contain tens of lemmas), and the outcome should be quick +enough. +" + +attribute [mfld_simps] id.def function.comp.left_id set.mem_set_of_eq and_true true_and + function.comp_app and_self eq_self_iff_true function.comp.right_id not_false_iff true_or or_true + +mk_simp_attribute monad_norm none with functor_norm + +mk_simp_attribute nontriviality "Simp lemmas for `nontriviality` tactic" + +mk_simp_attribute parity_simps "Simp attribute for lemmas about `even`" + +mk_simp_attribute push_cast "The `push_cast` simp attribute uses `norm_cast` lemmas +to move casts toward the leaf nodes of the expression." + +mk_simp_attribute split_if_reduction + "Simp set for if-then-else statements, used in the `split_ifs` tactic" + +attribute [split_if_reduction] if_pos if_neg dif_pos dif_neg if_congr + +mk_simp_attribute transport_simps +"The simpset `transport_simps` is used by the tactic `transport` +to simplify certain expressions involving application of equivalences, +and trivial `eq.rec` or `ep.mpr` conversions. +It's probably best not to adjust it without understanding the algorithm used by `transport`." + +attribute [transport_simps] cast_eq + +mk_simp_attribute typevec "simp set for the manipulation of typevec and arrow expressions" diff --git a/src/tactic/norm_cast.lean b/src/tactic/norm_cast.lean index dd64e04597622..b0c4fdc0f9409 100644 --- a/src/tactic/norm_cast.lean +++ b/src/tactic/norm_cast.lean @@ -61,9 +61,6 @@ when_tracing `norm_cast $ do a ← pp a, trace ("[norm_cast] " ++ msg ++ a : format) -mk_simp_attribute push_cast "The `push_cast` simp attribute uses `norm_cast` lemmas -to move casts toward the leaf nodes of the expression." - /-- `label` is a type used to classify `norm_cast` lemmas. * elim lemma: LHS has 0 head coes and ≥ 1 internal coe diff --git a/src/tactic/split_ifs.lean b/src/tactic/split_ifs.lean index dd06561183e13..e96604ceea6a0 100644 --- a/src/tactic/split_ifs.lean +++ b/src/tactic/split_ifs.lean @@ -27,11 +27,6 @@ lctx ← at_.get_locals, lctx ← lctx.mmap infer_type, tgt ← target, let es := if at_.include_goal then tgt::lctx else lctx, return $ find_if_cond $ es.foldr app default -run_cmd mk_simp_attr `split_if_reduction -run_cmd add_doc_string `simp_attr.split_if_reduction "Simp set for if-then-else statements" - -attribute [split_if_reduction] if_pos if_neg dif_pos dif_neg if_congr - meta def reduce_ifs_at (at_ : loc) : tactic unit := do sls ← get_user_simp_lemmas `split_if_reduction, let cfg : simp_config := { fail_if_unchanged := ff }, diff --git a/src/tactic/transport.lean b/src/tactic/transport.lean index 2e7ec3dc4a02d..53bdde01db08e 100644 --- a/src/tactic/transport.lean +++ b/src/tactic/transport.lean @@ -19,23 +19,6 @@ to a `monoid β`, the new multiplication is definitionally `λ x y, e (e.symm a namespace tactic open tactic.interactive -mk_simp_attribute transport_simps -"The simpset `transport_simps` is used by the tactic `transport` -to simplify certain expressions involving application of equivalences, -and trivial `eq.rec` or `ep.mpr` conversions. -It's probably best not to adjust it without understanding the algorithm used by `transport`." - -attribute [transport_simps] - eq_rec_constant - eq_mp_eq_cast - cast_eq - equiv.to_fun_as_coe - equiv.arrow_congr'_apply - equiv.symm_apply_apply - -- we use `apply_eq_iff_eq_symm_apply` rather than `apply_eq_iff_eq`, - -- as many axioms have a constant on the right-hand-side - equiv.apply_eq_iff_eq_symm_apply - /-- Given `s : S α` for some structure `S` depending on a type `α`, and an equivalence `e : α ≃ β`, diff --git a/test/equiv_rw.lean b/test/equiv_rw.lean index a39015992dc92..ae548862f1509 100644 --- a/test/equiv_rw.lean +++ b/test/equiv_rw.lean @@ -276,9 +276,9 @@ end -- The constructions and proofs here are written as uniformly as possible. -- This example is the blueprint for the `transport` tactic. -mk_simp_attribute transport_simps "simps useful inside `transport`" +mk_simp_attribute transport_simps' "simps useful inside `transport`" -attribute [transport_simps] +attribute [transport_simps'] eq_rec_constant cast_eq equiv.to_fun_as_coe @@ -291,7 +291,7 @@ begin refine_struct { .. }, { have mul := S.mul, equiv_rw e at mul, exact mul, }, { try { unfold_projs }, - simp only with transport_simps, + simp only with transport_simps', have mul_assoc := S.mul_assoc, equiv_rw e at mul_assoc, solve_by_elim, }, From d0b1936853671209a866fa35b9e54949c81116e2 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Fri, 30 Jun 2023 15:09:19 +0000 Subject: [PATCH 1230/1291] chore(*): add mathlib4 synchronization comments (#19220) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.category.fgModule.limits` * `algebra.homology.differential_object` * `algebraic_geometry.AffineScheme` * `algebraic_geometry.gluing` * `algebraic_geometry.limits` * `algebraic_geometry.properties` * `algebraic_geometry.pullbacks` * `linear_algebra.clifford_algebra.even` * `number_theory.legendre_symbol.norm_num` * `number_theory.zeta_function` * `order.category.omega_complete_partial_order` * `ring_theory.witt_vector.frobenius_fraction_field` * `set_theory.game.nim` * `set_theory.surreal.dyadic` * `topology.sheaves.skyscraper` --- src/algebra/category/fgModule/limits.lean | 3 +++ src/algebra/homology/differential_object.lean | 3 +++ src/algebraic_geometry/AffineScheme.lean | 3 +++ src/algebraic_geometry/gluing.lean | 3 +++ src/algebraic_geometry/limits.lean | 3 +++ src/algebraic_geometry/properties.lean | 3 +++ src/algebraic_geometry/pullbacks.lean | 3 +++ src/linear_algebra/clifford_algebra/even.lean | 3 +++ src/number_theory/legendre_symbol/norm_num.lean | 3 +++ src/number_theory/zeta_function.lean | 3 +++ src/order/category/omega_complete_partial_order.lean | 3 +++ src/ring_theory/witt_vector/frobenius_fraction_field.lean | 3 +++ src/set_theory/game/nim.lean | 3 +++ src/set_theory/surreal/dyadic.lean | 3 +++ src/topology/sheaves/skyscraper.lean | 3 +++ 15 files changed, 45 insertions(+) diff --git a/src/algebra/category/fgModule/limits.lean b/src/algebra/category/fgModule/limits.lean index f4c114b7dccdb..6c02caa86c100 100644 --- a/src/algebra/category/fgModule/limits.lean +++ b/src/algebra/category/fgModule/limits.lean @@ -14,6 +14,9 @@ import category_theory.limits.constructions.limits_of_products_and_equalizers /-! # `forget₂ (fgModule K) (Module K)` creates all finite limits. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + And hence `fgModule K` has all finite limits. ## Future work diff --git a/src/algebra/homology/differential_object.lean b/src/algebra/homology/differential_object.lean index 12a6920cb0ba9..03f8f3839d349 100644 --- a/src/algebra/homology/differential_object.lean +++ b/src/algebra/homology/differential_object.lean @@ -9,6 +9,9 @@ import category_theory.differential_object /-! # Homological complexes are differential graded objects. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We verify that a `homological_complex` indexed by an `add_comm_group` is essentially the same thing as a differential graded object. diff --git a/src/algebraic_geometry/AffineScheme.lean b/src/algebraic_geometry/AffineScheme.lean index 4def63a940a0a..827f2d0f8852b 100644 --- a/src/algebraic_geometry/AffineScheme.lean +++ b/src/algebraic_geometry/AffineScheme.lean @@ -11,6 +11,9 @@ import ring_theory.localization.inv_submonoid /-! # Affine schemes +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define the category of `AffineScheme`s as the essential image of `Spec`. We also define predicates about affine schemes and affine open sets. diff --git a/src/algebraic_geometry/gluing.lean b/src/algebraic_geometry/gluing.lean index 2a5345bbf6a65..15c3b25257934 100644 --- a/src/algebraic_geometry/gluing.lean +++ b/src/algebraic_geometry/gluing.lean @@ -9,6 +9,9 @@ import algebraic_geometry.open_immersion.Scheme /-! # Gluing Schemes +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Given a family of gluing data of schemes, we may glue them together. ## Main definitions diff --git a/src/algebraic_geometry/limits.lean b/src/algebraic_geometry/limits.lean index 6268fba92246e..22c9fbc92cf7f 100644 --- a/src/algebraic_geometry/limits.lean +++ b/src/algebraic_geometry/limits.lean @@ -9,6 +9,9 @@ import algebraic_geometry.AffineScheme /-! # (Co)Limits of Schemes +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We construct various limits and colimits in the category of schemes. * The existence of fibred products was shown in `algebraic_geometry/pullbacks.lean`. diff --git a/src/algebraic_geometry/properties.lean b/src/algebraic_geometry/properties.lean index 62429e066242e..7dbae1c65bc70 100644 --- a/src/algebraic_geometry/properties.lean +++ b/src/algebraic_geometry/properties.lean @@ -12,6 +12,9 @@ import ring_theory.local_properties /-! # Basic properties of schemes +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We provide some basic properties of schemes ## Main definition diff --git a/src/algebraic_geometry/pullbacks.lean b/src/algebraic_geometry/pullbacks.lean index 726ffd94eed13..5173a80bbca4d 100644 --- a/src/algebraic_geometry/pullbacks.lean +++ b/src/algebraic_geometry/pullbacks.lean @@ -11,6 +11,9 @@ import category_theory.limits.shapes.diagonal /-! # Fibred products of schemes +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we construct the fibred product of schemes via gluing. We roughly follow [har77] Theorem 3.3. diff --git a/src/linear_algebra/clifford_algebra/even.lean b/src/linear_algebra/clifford_algebra/even.lean index e977571460879..a374a0ead8693 100644 --- a/src/linear_algebra/clifford_algebra/even.lean +++ b/src/linear_algebra/clifford_algebra/even.lean @@ -9,6 +9,9 @@ import linear_algebra.clifford_algebra.grading /-! # The universal property of the even subalgebra +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions * `clifford_algebra.even Q`: The even subalgebra of `clifford_algebra Q`. diff --git a/src/number_theory/legendre_symbol/norm_num.lean b/src/number_theory/legendre_symbol/norm_num.lean index 00cfe5eb881ff..a47ded8453ec7 100644 --- a/src/number_theory/legendre_symbol/norm_num.lean +++ b/src/number_theory/legendre_symbol/norm_num.lean @@ -8,6 +8,9 @@ import number_theory.legendre_symbol.jacobi_symbol /-! # A `norm_num` extension for Jacobi and Legendre symbols +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We extend the `tactic.interactive.norm_num` tactic so that it can be used to provably compute the value of the Jacobi symbol `J(a | b)` or the Legendre symbol `legendre_sym p a` when the arguments are numerals. diff --git a/src/number_theory/zeta_function.lean b/src/number_theory/zeta_function.lean index b3dc94155d253..cda925a279f9f 100644 --- a/src/number_theory/zeta_function.lean +++ b/src/number_theory/zeta_function.lean @@ -10,6 +10,9 @@ import number_theory.zeta_values /-! # Definition of the Riemann zeta function +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions: * `riemann_zeta`: the Riemann zeta function `ζ : ℂ → ℂ`. diff --git a/src/order/category/omega_complete_partial_order.lean b/src/order/category/omega_complete_partial_order.lean index 65d74186c9172..15b627c09c367 100644 --- a/src/order/category/omega_complete_partial_order.lean +++ b/src/order/category/omega_complete_partial_order.lean @@ -13,6 +13,9 @@ import category_theory.concrete_category.bundled_hom /-! # Category of types with a omega complete partial order +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we bundle the class `omega_complete_partial_order` into a concrete category and prove that continuous functions also form a `omega_complete_partial_order`. diff --git a/src/ring_theory/witt_vector/frobenius_fraction_field.lean b/src/ring_theory/witt_vector/frobenius_fraction_field.lean index d251d81d5ad0b..e9cd65c73af4f 100644 --- a/src/ring_theory/witt_vector/frobenius_fraction_field.lean +++ b/src/ring_theory/witt_vector/frobenius_fraction_field.lean @@ -10,6 +10,9 @@ import ring_theory.witt_vector.discrete_valuation_ring /-! # Solving equations about the Frobenius map on the field of fractions of `𝕎 k` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The goal of this file is to prove `witt_vector.exists_frobenius_solution_fraction_ring`, which says that for an algebraically closed field `k` of characteristic `p` and `a, b` in the field of fractions of Witt vectors over `k`, diff --git a/src/set_theory/game/nim.lean b/src/set_theory/game/nim.lean index fbbbb91d3abbd..93d6b3f894ccd 100644 --- a/src/set_theory/game/nim.lean +++ b/src/set_theory/game/nim.lean @@ -10,6 +10,9 @@ import set_theory.game.impartial /-! # Nim and the Sprague-Grundy theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains the definition for nim for any ordinal `o`. In the game of `nim o₁` both players may move to `nim o₂` for any `o₂ < o₁`. We also define a Grundy value for an impartial game `G` and prove the Sprague-Grundy theorem, that diff --git a/src/set_theory/surreal/dyadic.lean b/src/set_theory/surreal/dyadic.lean index e663ebaecd75e..0f666a31e2705 100644 --- a/src/set_theory/surreal/dyadic.lean +++ b/src/set_theory/surreal/dyadic.lean @@ -10,6 +10,9 @@ import ring_theory.localization.basic /-! # Dyadic numbers + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. Dyadic numbers are obtained by localizing ℤ away from 2. They are the initial object in the category of rings with no 2-torsion. diff --git a/src/topology/sheaves/skyscraper.lean b/src/topology/sheaves/skyscraper.lean index c02f854b23030..08b19fe467af6 100644 --- a/src/topology/sheaves/skyscraper.lean +++ b/src/topology/sheaves/skyscraper.lean @@ -10,6 +10,9 @@ import topology.sheaves.functors /-! # Skyscraper (pre)sheaves +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A skyscraper (pre)sheaf `𝓕 : (pre)sheaf C X` is the (pre)sheaf with value `A` at point `p₀` that is supported only at open sets contain `p₀`, i.e. `𝓕(U) = A` if `p₀ ∈ U` and `𝓕(U) = *` if `p₀ ∉ U` where `*` is a terminal object of `C`. In terms of stalks, `𝓕` is supported at all specializations From 44b3f42373ba4f6daed7c340d262112db2468e97 Mon Sep 17 00:00:00 2001 From: leanprover-community-bot Date: Sat, 1 Jul 2023 04:27:34 +0000 Subject: [PATCH 1231/1291] chore(scripts): update nolints.txt (#19227) I am happy to remove some nolints for you! --- scripts/style-exceptions.txt | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/scripts/style-exceptions.txt b/scripts/style-exceptions.txt index c8f9e355af435..c04c94940278a 100644 --- a/scripts/style-exceptions.txt +++ b/scripts/style-exceptions.txt @@ -1,4 +1,4 @@ -src/control/basic.lean : line 9 : ERR_MOD : Module docstring missing, or too late +src/control/basic.lean : line 10 : ERR_MOD : Module docstring missing, or too late src/control/monad/cont.lean : line 13 : ERR_MOD : Module docstring missing, or too late src/control/monad/writer.lean : line 11 : ERR_MOD : Module docstring missing, or too late src/control/traversable/derive.lean : line 11 : ERR_MOD : Module docstring missing, or too late From a176cb1219e300e85793d44583dede42377b51af Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sun, 2 Jul 2023 06:51:23 +0000 Subject: [PATCH 1232/1291] feat(linear_algebra/matrix/schur_complement): invertibility of block matrices (#19156) Co-authored-by: Jeremy Tan Jie Rui Co-authored-by: Jireh Loreaux --- .../preadditive/biproducts.lean | 7 + .../matrix/schur_complement.lean | 153 ++++++++++++++++++ 2 files changed, 160 insertions(+) diff --git a/src/category_theory/preadditive/biproducts.lean b/src/category_theory/preadditive/biproducts.lean index b6f15e50a9d78..e50ebb7b168d3 100644 --- a/src/category_theory/preadditive/biproducts.lean +++ b/src/category_theory/preadditive/biproducts.lean @@ -46,6 +46,13 @@ In (or between) preadditive categories, * A functor preserves a biproduct if and only if it preserves the corresponding product if and only if it preserves the corresponding coproduct. + +There are connections between this material and the special case of the category whose morphisms are +matrices over a ring, in particular the Schur complement (see +`linear_algebra.matrix.schur_complement`). In particular, the declarations +`category_theory.biprod.iso_elim`, `category_theory.biprod.gaussian` +and `matrix.invertible_of_from_blocks₁₁_invertible` are all closely related. + -/ open category_theory diff --git a/src/linear_algebra/matrix/schur_complement.lean b/src/linear_algebra/matrix/schur_complement.lean index 440f1fbe308be..1c354e0f09131 100644 --- a/src/linear_algebra/matrix/schur_complement.lean +++ b/src/linear_algebra/matrix/schur_complement.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 Alexander Bentkamp. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Alexander Bentkamp, Eric Wieser, Jeremy Avigad, Johan Commelin -/ +import data.matrix.invertible import linear_algebra.matrix.nonsingular_inverse import linear_algebra.matrix.pos_def @@ -11,6 +12,11 @@ import linear_algebra.matrix.pos_def This file proves properties of 2×2 block matrices `[A B; C D]` that relate to the Schur complement `D - C⬝A⁻¹⬝B`. +Some of the results here generalize to 2×2 matrices in a category, rather than just a ring. A few +results in this direction can be found in the the file `cateogry_theory.preadditive.biproducts`, +especially the declarations `category_theory.biprod.gaussian` and `category_theory.biprod.iso_elim`. +Compare with `matrix.invertible_of_from_blocks₁₁_invertible`. + ## Main results * `matrix.det_from_blocks₁₁`, `matrix.det_from_blocks₂₂`: determinant of a block matrix in terms of @@ -215,6 +221,153 @@ end end triangular +/-! ### 2×2 block matrices -/ + +section block + +/-! #### General 2×2 block matrices-/ + +/-- A block matrix is invertible if the bottom right corner and the corresponding schur complement +is. -/ +def from_blocks₂₂_invertible + (A : matrix m m α) (B : matrix m n α) (C : matrix n m α) (D : matrix n n α) + [invertible D] [invertible (A - B⬝⅟D⬝C)] : + invertible (from_blocks A B C D) := +begin + -- factor `from_blocks` via `from_blocks_eq_of_invertible₂₂`, and state the inverse we expect + refine invertible.copy' _ _ + (from_blocks + (⅟(A - B⬝⅟D⬝C)) (-(⅟(A - B⬝⅟D⬝C)⬝B⬝⅟D)) + (-(⅟D⬝C⬝⅟(A - B⬝⅟D⬝C))) (⅟D + ⅟D⬝C⬝⅟(A - B⬝⅟D⬝C)⬝B⬝⅟D)) + (from_blocks_eq_of_invertible₂₂ _ _ _ _) _, + { -- the product is invertible because all the factors are + letI : invertible (1 : matrix n n α) := invertible_one, + letI : invertible (1 : matrix m m α) := invertible_one, + refine invertible.matrix_mul _ (from_blocks_zero₁₂_invertible _ _ _), + exact invertible.matrix_mul (from_blocks_zero₂₁_invertible _ _ _) + (from_blocks_zero₂₁_invertible _ _ _) }, + { -- unfold the `invertible` instances to get the raw factors + show _ = from_blocks 1 0 (-(1 ⬝ (⅟ D ⬝ C) ⬝ 1)) 1 + ⬝ (from_blocks (⅟ (A - B ⬝ ⅟ D ⬝ C)) (-(⅟ (A - B ⬝ ⅟ D ⬝ C) ⬝ 0 ⬝ ⅟ D)) 0 (⅟ D) + ⬝ from_blocks 1 (-(1 ⬝ (B ⬝ ⅟ D) ⬝ 1)) 0 1), + -- combine into a single block matrix + simp only [from_blocks_multiply, inv_of_one, matrix.one_mul, matrix.mul_one, matrix.zero_mul, + matrix.mul_zero, add_zero, zero_add, neg_zero, matrix.mul_neg, matrix.neg_mul, neg_neg, + ←matrix.mul_assoc, add_comm], }, +end + +/-- A block matrix is invertible if the top left corner and the corresponding schur complement +is. -/ +def from_blocks₁₁_invertible + (A : matrix m m α) (B : matrix m n α) (C : matrix n m α) (D : matrix n n α) + [invertible A] [invertible (D - C⬝⅟A⬝B)] : + invertible (from_blocks A B C D) := +begin + -- we argue by symmetry + letI := from_blocks₂₂_invertible D C B A, + letI iDCBA + := + submatrix_equiv_invertible (from_blocks D C B A) (equiv.sum_comm _ _) (equiv.sum_comm _ _), + exact iDCBA.copy' _ + (from_blocks + (⅟A + ⅟A⬝B⬝⅟(D - C⬝⅟A⬝B)⬝C⬝⅟A) (-(⅟A⬝B⬝⅟(D - C⬝⅟A⬝B))) + (-(⅟(D - C⬝⅟A⬝B)⬝C⬝⅟A)) (⅟(D - C⬝⅟A⬝B))) + (from_blocks_submatrix_sum_swap_sum_swap _ _ _ _).symm + (from_blocks_submatrix_sum_swap_sum_swap _ _ _ _).symm, +end + +lemma inv_of_from_blocks₂₂_eq + (A : matrix m m α) (B : matrix m n α) (C : matrix n m α) (D : matrix n n α) + [invertible D] [invertible (A - B⬝⅟D⬝C)] [invertible (from_blocks A B C D)] : + ⅟(from_blocks A B C D) = from_blocks + (⅟(A - B⬝⅟D⬝C)) (-(⅟(A - B⬝⅟D⬝C)⬝B⬝⅟D)) + (-(⅟D⬝C⬝⅟(A - B⬝⅟D⬝C))) (⅟D + ⅟D⬝C⬝⅟(A - B⬝⅟D⬝C)⬝B⬝⅟D):= +begin + letI := from_blocks₂₂_invertible A B C D, + convert (rfl : ⅟(from_blocks A B C D) = _), +end + +lemma inv_of_from_blocks₁₁_eq + (A : matrix m m α) (B : matrix m n α) (C : matrix n m α) (D : matrix n n α) + [invertible A] [invertible (D - C⬝⅟A⬝B)] [invertible (from_blocks A B C D)] : + ⅟(from_blocks A B C D) = from_blocks + (⅟A + ⅟A⬝B⬝⅟(D - C⬝⅟A⬝B)⬝C⬝⅟A) (-(⅟A⬝B⬝⅟(D - C⬝⅟A⬝B))) + (-(⅟(D - C⬝⅟A⬝B)⬝C⬝⅟A)) (⅟(D - C⬝⅟A⬝B)) := +begin + letI := from_blocks₁₁_invertible A B C D, + convert (rfl : ⅟(from_blocks A B C D) = _), +end + +/-- If a block matrix is invertible and so is its bottom left element, then so is the corresponding +Schur complement. -/ +def invertible_of_from_blocks₂₂_invertible + (A : matrix m m α) (B : matrix m n α) (C : matrix n m α) (D : matrix n n α) + [invertible D] [invertible (from_blocks A B C D)] : invertible (A - B⬝⅟D⬝C) := +begin + suffices : invertible (from_blocks (A - B ⬝ ⅟ D ⬝ C) 0 0 D), + { exactI (invertible_of_from_blocks_zero₁₂_invertible (A - B ⬝ ⅟ D ⬝ C) 0 D).1 }, + letI : invertible (1 : matrix n n α) := invertible_one, + letI : invertible (1 : matrix m m α) := invertible_one, + letI iDC : invertible (from_blocks 1 0 (⅟ D ⬝ C) 1 : matrix (m ⊕ n) (m ⊕ n) α) := + from_blocks_zero₁₂_invertible _ _ _, + letI iBD : invertible (from_blocks 1 (B ⬝ ⅟ D) 0 1 : matrix(m ⊕ n) (m ⊕ n) α) := + from_blocks_zero₂₁_invertible _ _ _, + letI iBDC := invertible.copy ‹_› _ (from_blocks_eq_of_invertible₂₂ A B C D).symm, + refine (iBD.matrix_mul_left _).symm _, + refine (iDC.matrix_mul_right _).symm iBDC, +end + +/-- If a block matrix is invertible and so is its bottom left element, then so is the corresponding +Schur complement. -/ +def invertible_of_from_blocks₁₁_invertible + (A : matrix m m α) (B : matrix m n α) (C : matrix n m α) (D : matrix n n α) + [invertible A] [invertible (from_blocks A B C D)] : invertible (D - C⬝⅟A⬝B) := +begin + -- another symmetry argument + letI iABCD' := + submatrix_equiv_invertible (from_blocks A B C D) (equiv.sum_comm _ _) (equiv.sum_comm _ _), + letI iDCBA := iABCD'.copy _ (from_blocks_submatrix_sum_swap_sum_swap _ _ _ _).symm, + refine invertible_of_from_blocks₂₂_invertible D C B A, +end + +/-- `matrix.invertible_of_from_blocks₂₂_invertible` and `matrix.from_blocks₂₂_invertible` as an +equivalence. -/ +def invertible_equiv_from_blocks₂₂_invertible + (A : matrix m m α) (B : matrix m n α) (C : matrix n m α) (D : matrix n n α) + [invertible D] : invertible (from_blocks A B C D) ≃ invertible (A - B⬝⅟D⬝C) := +{ to_fun := λ iABCD, by exactI invertible_of_from_blocks₂₂_invertible _ _ _ _, + inv_fun := λ i_schur,by exactI from_blocks₂₂_invertible _ _ _ _, + left_inv := λ iABCD, subsingleton.elim _ _, + right_inv := λ i_schur, subsingleton.elim _ _ } + +/-- `matrix.invertible_of_from_blocks₁₁_invertible` and `matrix.from_blocks₁₁_invertible` as an +equivalence. -/ +def invertible_equiv_from_blocks₁₁_invertible + (A : matrix m m α) (B : matrix m n α) (C : matrix n m α) (D : matrix n n α) + [invertible A] : invertible (from_blocks A B C D) ≃ invertible (D - C⬝⅟A⬝B) := +{ to_fun := λ iABCD, by exactI invertible_of_from_blocks₁₁_invertible _ _ _ _, + inv_fun := λ i_schur,by exactI from_blocks₁₁_invertible _ _ _ _, + left_inv := λ iABCD, subsingleton.elim _ _, + right_inv := λ i_schur, subsingleton.elim _ _ } + +/-- If the bottom-left element of a block matrix is invertible, then the whole matrix is invertible +iff the corresponding schur complement is. -/ +lemma is_unit_from_blocks_iff_of_invertible₂₂ + {A : matrix m m α} {B : matrix m n α} {C : matrix n m α} {D : matrix n n α} [invertible D] : + is_unit (from_blocks A B C D) ↔ is_unit (A - B⬝⅟D⬝C) := +by simp only [← nonempty_invertible_iff_is_unit, + (invertible_equiv_from_blocks₂₂_invertible A B C D).nonempty_congr] + +/-- If the top-right element of a block matrix is invertible, then the whole matrix is invertible +iff the corresponding schur complement is. -/ +lemma is_unit_from_blocks_iff_of_invertible₁₁ + {A : matrix m m α} {B : matrix m n α} {C : matrix n m α} {D : matrix n n α} [invertible A] : + is_unit (from_blocks A B C D) ↔ is_unit (D - C⬝⅟A⬝B) := +by simp only [← nonempty_invertible_iff_is_unit, + (invertible_equiv_from_blocks₁₁_invertible A B C D).nonempty_congr] + +end block + /-! ### Lemmas about `matrix.det` -/ section det From e473c3198bb41f68560cab68a0529c854b618833 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sun, 2 Jul 2023 14:35:26 +0000 Subject: [PATCH 1233/1291] refactor: redefine `bundle.total_space` (#19221) - Use a custom structure for `bundle.total_space`. - Use `bundle.total_space.mk` instead of `bundle.total_space_mk`. - Use `bundle.total_space.to_prod` instead of `equiv.sigma_equiv_prod`. - Use `bundle.total_space.mk'` (scoped notation) to specify `F`. - Rename `bundle.trivial.proj_snd` to `bundle.total_space.trivial_snd` to allow dot notation. Should we just use `bundle.total_space.snd` since `bundle.trivial` is now reducible? - Add an unused argument to `bundle.total_space`. - Make `bundle.trivial` and `bundle.continuous_linear_map` reducible. - Drop instances that are no longer needed. --- src/data/bundle.lean | 126 +++++----- src/geometry/manifold/cont_mdiff_mfderiv.lean | 47 ++-- src/geometry/manifold/mfderiv.lean | 49 ++-- .../manifold/vector_bundle/basic.lean | 81 ++++--- src/geometry/manifold/vector_bundle/hom.lean | 14 +- .../manifold/vector_bundle/pullback.lean | 2 +- .../vector_bundle/smooth_section.lean | 18 +- .../manifold/vector_bundle/tangent.lean | 18 +- src/topology/covering.lean | 4 +- src/topology/fiber_bundle/basic.lean | 131 +++++----- src/topology/fiber_bundle/constructions.lean | 135 ++++++----- src/topology/fiber_bundle/trivialization.lean | 73 +++--- src/topology/vector_bundle/basic.lean | 223 +++++++++--------- src/topology/vector_bundle/constructions.lean | 16 +- src/topology/vector_bundle/hom.lean | 112 ++++----- 15 files changed, 501 insertions(+), 548 deletions(-) diff --git a/src/data/bundle.lean b/src/data/bundle.lean index 68cfd8bbb915b..b916ce0d12457 100644 --- a/src/data/bundle.lean +++ b/src/data/bundle.lean @@ -15,15 +15,26 @@ should contain all possible results that do not involve any topology. We represent a bundle `E` over a base space `B` as a dependent type `E : B → Type*`. -We provide a type synonym of `Σ x, E x` as `bundle.total_space E`, to be able to endow it with -a topology which is not the disjoint union topology `sigma.topological_space`. In general, the -constructions of fiber bundles we will make will be of this form. +We define `bundle.total_space F E` to be the type of pairs `⟨b, x⟩`, where `b : B` and `x : E x`. +This type is isomorphic to `Σ x, E x` and uses an extra argument `F` for reasons explained below. In +general, the constructions of fiber bundles we will make will be of this form. ## Main Definitions * `bundle.total_space` the total space of a bundle. * `bundle.total_space.proj` the projection from the total space to the base space. -* `bundle.total_space_mk` the constructor for the total space. +* `bundle.total_space.mk` the constructor for the total space. + +## Implementation Notes + +- We use a custom structure for the total space of a bundle instead of using a type synonym for the + canonical disjoint union `Σ x, E x` because the total space usually has a different topology and + Lean 4 `simp` fails to apply lemmas about `Σ x, E x` to elements of the total space. + +- The definition of `bundle.total_space` has an unused argument `F`. The reason is that in some + constructions (e.g., `bundle.continuous_linear_map.vector_bundle`) we need access to the atlas of + trivializations of original fiber bundles to construct the topology on the total space of the new + fiber bundle. ## References - https://en.wikipedia.org/wiki/Bundle_(mathematics) @@ -31,61 +42,65 @@ constructions of fiber bundles we will make will be of this form. namespace bundle -variables {B : Type*} (E : B → Type*) +variables {B F : Type*} (E : B → Type*) /-- -`bundle.total_space E` is the total space of the bundle `Σ x, E x`. -This type synonym is used to avoid conflicts with general sigma types. +`bundle.total_space E` is the total space of the bundle. It consists of pairs +`(proj : B, snd : E proj)`. -/ -def total_space := Σ x, E x +@[ext] +structure total_space (F : Type*) (E : B → Type*) := +(proj : B) +(snd : E proj) instance [inhabited B] [inhabited (E default)] : - inhabited (total_space E) := ⟨⟨default, default⟩⟩ + inhabited (total_space F E) := ⟨⟨default, default⟩⟩ variables {E} /-- `bundle.total_space.proj` is the canonical projection `bundle.total_space E → B` from the total space to the base space. -/ -@[simp, reducible] def total_space.proj : total_space E → B := sigma.fst +add_decl_doc total_space.proj -- this notation won't be used in the pretty-printer localized "notation `π` := @bundle.total_space.proj _" in bundle -/-- Constructor for the total space of a bundle. -/ -@[simp, reducible] def total_space_mk (b : B) (a : E b) : - bundle.total_space E := ⟨b, a⟩ - -lemma total_space.proj_mk {x : B} {y : E x} : (total_space_mk x y).proj = x := -rfl - -lemma sigma_mk_eq_total_space_mk {x : B} {y : E x} : sigma.mk x y = total_space_mk x y := -rfl +-- TODO: try `abbrev` in Lean 4 +localized "notation `total_space.mk'` F:max := @bundle.total_space.mk _ F _" in bundle lemma total_space.mk_cast {x x' : B} (h : x = x') (b : E x) : - total_space_mk x' (cast (congr_arg E h) b) = total_space_mk x b := + total_space.mk' F x' (cast (congr_arg E h) b) = total_space.mk x b := by { subst h, refl } -lemma total_space.eta (z : total_space E) : - total_space_mk z.proj z.2 = z := -sigma.eta z +instance {x : B} : has_coe_t (E x) (total_space F E) := ⟨total_space.mk x⟩ -instance {x : B} : has_coe_t (E x) (total_space E) := ⟨total_space_mk x⟩ +@[simp] lemma total_space.coe_proj (x : B) (v : E x) : (v : total_space F E).proj = x := rfl +@[simp] lemma total_space.coe_snd {x : B} {y : E x} : (y : total_space F E).snd = y := rfl -@[simp] lemma coe_fst (x : B) (v : E x) : (v : total_space E).fst = x := rfl -@[simp] lemma coe_snd {x : B} {y : E x} : (y : total_space E).snd = y := rfl +lemma total_space.coe_eq_mk {x : B} (v : E x) : (v : total_space F E) = total_space.mk x v := rfl -lemma to_total_space_coe {x : B} (v : E x) : (v : total_space E) = total_space_mk x v := rfl +lemma total_space.eta (z : total_space F E) : + total_space.mk z.proj z.2 = z := +by cases z; refl -- notation for the direct sum of two bundles over the same base notation E₁ ` ×ᵇ `:100 E₂ := λ x, E₁ x × E₂ x /-- `bundle.trivial B F` is the trivial bundle over `B` of fiber `F`. -/ -def trivial (B : Type*) (F : Type*) : B → Type* := function.const B F - -instance {F : Type*} [inhabited F] {b : B} : inhabited (bundle.trivial B F b) := ⟨(default : F)⟩ +@[reducible, nolint unused_arguments] +def trivial (B : Type*) (F : Type*) : B → Type* := λ _, F /-- The trivial bundle, unlike other bundles, has a canonical projection on the fiber. -/ -def trivial.proj_snd (B : Type*) (F : Type*) : total_space (bundle.trivial B F) → F := sigma.snd +def total_space.trivial_snd (B : Type*) (F : Type*) : total_space F (bundle.trivial B F) → F := +total_space.snd + +/-- A trivial bundle is equivalent to the product `B × F`. -/ +@[simps { attrs := [`simp, `mfld_simps] }] +def total_space.to_prod (B F : Type*) : total_space F (λ _ : B, F) ≃ B × F := +{ to_fun := λ x, (x.1, x.2), + inv_fun := λ x, ⟨x.1, x.2⟩, + left_inv := λ ⟨_, _⟩, rfl, + right_inv := λ ⟨_, _⟩, rfl } section pullback @@ -93,55 +108,36 @@ variable {B' : Type*} /-- The pullback of a bundle `E` over a base `B` under a map `f : B' → B`, denoted by `pullback f E` or `f *ᵖ E`, is the bundle over `B'` whose fiber over `b'` is `E (f b')`. -/ -@[nolint has_nonempty_instance] def pullback (f : B' → B) (E : B → Type*) := λ x, E (f x) +def pullback (f : B' → B) (E : B → Type*) : B' → Type* := λ x, E (f x) + +notation f ` *ᵖ ` E:max := pullback f E -notation f ` *ᵖ ` E := pullback f E +instance {f : B' → B} {x : B'} [nonempty (E (f x))] : nonempty (f *ᵖ E x) := ‹nonempty (E (f x))› /-- Natural embedding of the total space of `f *ᵖ E` into `B' × total_space E`. -/ @[simp] def pullback_total_space_embedding (f : B' → B) : - total_space (f *ᵖ E) → B' × total_space E := -λ z, (z.proj, total_space_mk (f z.proj) z.2) + total_space F (f *ᵖ E) → B' × total_space F E := +λ z, (z.proj, total_space.mk (f z.proj) z.2) /-- The base map `f : B' → B` lifts to a canonical map on the total spaces. -/ -def pullback.lift (f : B' → B) : total_space (f *ᵖ E) → total_space E := -λ z, total_space_mk (f z.proj) z.2 +@[simps { attrs := [`simp, `mfld_simps] }] +def pullback.lift (f : B' → B) : total_space F (f *ᵖ E) → total_space F E := +λ z, ⟨f z.proj, z.2⟩ -@[simp] lemma pullback.proj_lift (f : B' → B) (x : total_space (f *ᵖ E)) : - (pullback.lift f x).proj = f x.1 := -rfl - -@[simp] lemma pullback.lift_mk (f : B' → B) (x : B') (y : E (f x)) : - pullback.lift f (total_space_mk x y) = total_space_mk (f x) y := -rfl - -lemma pullback_total_space_embedding_snd (f : B' → B) (x : total_space (f *ᵖ E)) : - (pullback_total_space_embedding f x).2 = pullback.lift f x := +@[simp, mfld_simps] lemma pullback.lift_mk (f : B' → B) (x : B') (y : E (f x)) : + pullback.lift f (total_space.mk' F x y) = ⟨f x, y⟩ := rfl end pullback section fiber_structures -variable [∀ x, add_comm_monoid (E x)] - -@[simp] lemma coe_snd_map_apply (x : B) (v w : E x) : - (↑(v + w) : total_space E).snd = (v : total_space E).snd + (w : total_space E).snd := rfl +@[simp] lemma coe_snd_map_apply [∀ x, has_add (E x)] (x : B) (v w : E x) : + (↑(v + w) : total_space F E).snd = (v : total_space F E).snd + (w : total_space F E).snd := rfl -variables (R : Type*) [semiring R] [∀ x, module R (E x)] - -@[simp] lemma coe_snd_map_smul (x : B) (r : R) (v : E x) : - (↑(r • v) : total_space E).snd = r • (v : total_space E).snd := rfl +@[simp] lemma coe_snd_map_smul {R} [∀ x, has_smul R (E x)] (x : B) (r : R) (v : E x) : + (↑(r • v) : total_space F E).snd = r • (v : total_space F E).snd := rfl end fiber_structures -section trivial_instances - -variables {F : Type*} {R : Type*} [semiring R] (b : B) - -instance [add_comm_monoid F] : add_comm_monoid (bundle.trivial B F b) := ‹add_comm_monoid F› -instance [add_comm_group F] : add_comm_group (bundle.trivial B F b) := ‹add_comm_group F› -instance [add_comm_monoid F] [module R F] : module R (bundle.trivial B F b) := ‹module R F› - -end trivial_instances - end bundle diff --git a/src/geometry/manifold/cont_mdiff_mfderiv.lean b/src/geometry/manifold/cont_mdiff_mfderiv.lean index 048675244325c..9efd7835a3bd5 100644 --- a/src/geometry/manifold/cont_mdiff_mfderiv.lean +++ b/src/geometry/manifold/cont_mdiff_mfderiv.lean @@ -248,7 +248,7 @@ space are model spaces in models with corners. The general fact is proved in lemma cont_mdiff_on.continuous_on_tangent_map_within_aux {f : H → H'} {s : set H} (hf : cont_mdiff_on I I' n f s) (hn : 1 ≤ n) (hs : unique_mdiff_on I s) : - continuous_on (tangent_map_within I I' f s) (π (tangent_space I) ⁻¹' s) := + continuous_on (tangent_map_within I I' f s) (π E (tangent_space I) ⁻¹' s) := begin suffices h : continuous_on (λ (p : H × E), (f p.fst, (fderiv_within 𝕜 (written_in_ext_chart_at I I' p.fst f) (I.symm ⁻¹' s ∩ range I) @@ -258,7 +258,7 @@ begin have B := ((tangent_bundle_model_space_homeomorph H' I').symm.continuous.comp_continuous_on h) .comp' A, have : (univ ∩ ⇑(tangent_bundle_model_space_homeomorph H I) ⁻¹' (prod.fst ⁻¹' s)) = - π (tangent_space I) ⁻¹' s, + π E (tangent_space I) ⁻¹' s, by { ext ⟨x, v⟩, simp only with mfld_simps }, rw this at B, apply B.congr, @@ -307,7 +307,8 @@ are model spaces in models with corners. The general fact is proved in lemma cont_mdiff_on.cont_mdiff_on_tangent_map_within_aux {f : H → H'} {s : set H} (hf : cont_mdiff_on I I' n f s) (hmn : m + 1 ≤ n) (hs : unique_mdiff_on I s) : - cont_mdiff_on I.tangent I'.tangent m (tangent_map_within I I' f s) (π (tangent_space I) ⁻¹' s) := + cont_mdiff_on I.tangent I'.tangent m (tangent_map_within I I' f s) + (π E (tangent_space I) ⁻¹' s) := begin have m_le_n : m ≤ n, { apply le_trans _ hmn, @@ -324,13 +325,13 @@ begin rw cont_mdiff_on_iff, refine ⟨hf.continuous_on_tangent_map_within_aux one_le_n hs, λp q, _⟩, have A : range I ×ˢ univ ∩ - ((equiv.sigma_equiv_prod H E).symm ∘ λ (p : E × E), ((I.symm) p.fst, p.snd)) ⁻¹' - (π (tangent_space I) ⁻¹' s) + ((total_space.to_prod H E).symm ∘ λ (p : E × E), ((I.symm) p.fst, p.snd)) ⁻¹' + (π E (tangent_space I) ⁻¹' s) = (range I ∩ I.symm ⁻¹' s) ×ˢ univ, - by { ext ⟨x, v⟩, simp only with mfld_simps }, + by { ext ⟨x, v⟩, simp only with mfld_simps, }, suffices h : cont_diff_on 𝕜 m (((λ (p : H' × E'), (I' p.fst, p.snd)) ∘ - (equiv.sigma_equiv_prod H' E')) ∘ tangent_map_within I I' f s ∘ - ((equiv.sigma_equiv_prod H E).symm) ∘ λ (p : E × E), (I.symm p.fst, p.snd)) + (total_space.to_prod H' E')) ∘ tangent_map_within I I' f s ∘ + ((total_space.to_prod H E).symm) ∘ λ (p : E × E), (I.symm p.fst, p.snd)) ((range ⇑I ∩ ⇑(I.symm) ⁻¹' s) ×ˢ univ), by simpa [A] using h, change cont_diff_on 𝕜 m (λ (p : E × E), @@ -369,7 +370,7 @@ is `C^m` when `m+1 ≤ n`. -/ theorem cont_mdiff_on.cont_mdiff_on_tangent_map_within (hf : cont_mdiff_on I I' n f s) (hmn : m + 1 ≤ n) (hs : unique_mdiff_on I s) : cont_mdiff_on I.tangent I'.tangent m (tangent_map_within I I' f s) - (π (tangent_space I) ⁻¹' s) := + (π E (tangent_space I) ⁻¹' s) := begin /- The strategy of the proof is to avoid unfolding the definitions, and reduce by functoriality to the case of functions on the model spaces, where we have already proved the result. @@ -407,20 +408,20 @@ begin let il := chart_at (model_prod H E) (tangent_map I I l p), let ir := chart_at (model_prod H' E') (tangent_map I I' (r ∘ f) p), let s' := f ⁻¹' r.source ∩ s ∩ l.source, - let s'_lift := π (tangent_space I) ⁻¹' s', + let s'_lift := π E (tangent_space I) ⁻¹' s', let s'l := l.target ∩ l.symm ⁻¹' s', - let s'l_lift := π (tangent_space I) ⁻¹' s'l, + let s'l_lift := π E (tangent_space I) ⁻¹' s'l, rcases continuous_on_iff'.1 hf'.1 r.source r.open_source with ⟨o, o_open, ho⟩, suffices h : cont_mdiff_on I.tangent I'.tangent m (tangent_map_within I I' f s) s'_lift, - { refine ⟨π (tangent_space I) ⁻¹' (o ∩ l.source), _, _, _⟩, - show is_open (π (tangent_space I) ⁻¹' (o ∩ l.source)), from + { refine ⟨π E (tangent_space I) ⁻¹' (o ∩ l.source), _, _, _⟩, + show is_open (π E (tangent_space I) ⁻¹' (o ∩ l.source)), from (is_open.inter o_open l.open_source).preimage (continuous_proj E _) , - show p ∈ π (tangent_space I) ⁻¹' (o ∩ l.source), + show p ∈ π E (tangent_space I) ⁻¹' (o ∩ l.source), { simp, have : p.proj ∈ f ⁻¹' r.source ∩ s, by simp [hp], rw ho at this, exact this.1 }, - { have : π (tangent_space I) ⁻¹' s ∩ π (tangent_space I) ⁻¹' (o ∩ l.source) = s'_lift, + { have : π E (tangent_space I) ⁻¹' s ∩ π E (tangent_space I) ⁻¹' (o ∩ l.source) = s'_lift, { dsimp only [s'_lift, s'], rw [ho], mfld_set_tac }, rw this, exact h } }, @@ -547,10 +548,10 @@ end derivative is continuous there. -/ theorem cont_mdiff_on.continuous_on_tangent_map_within (hf : cont_mdiff_on I I' n f s) (hmn : 1 ≤ n) (hs : unique_mdiff_on I s) : - continuous_on (tangent_map_within I I' f s) (π (tangent_space I) ⁻¹' s) := + continuous_on (tangent_map_within I I' f s) (π E (tangent_space I) ⁻¹' s) := begin have : cont_mdiff_on I.tangent I'.tangent 0 (tangent_map_within I I' f s) - (π (tangent_space I) ⁻¹' s) := + (π E (tangent_space I) ⁻¹' s) := hf.cont_mdiff_on_tangent_map_within hmn hs, exact this.continuous_on end @@ -599,15 +600,15 @@ may seem. TODO define splittings of vector bundles; state this result invariantly. -/ lemma tangent_map_tangent_bundle_pure (p : tangent_bundle I M) : - tangent_map I I.tangent (zero_section (tangent_space I)) p = ⟨⟨p.proj, 0⟩, ⟨p.2, 0⟩⟩ := + tangent_map I I.tangent (zero_section E (tangent_space I)) p = ⟨⟨p.proj, 0⟩, ⟨p.2, 0⟩⟩ := begin rcases p with ⟨x, v⟩, have N : I.symm ⁻¹' (chart_at H x).target ∈ 𝓝 (I ((chart_at H x) x)), { apply is_open.mem_nhds, apply (local_homeomorph.open_target _).preimage I.continuous_inv_fun, simp only with mfld_simps }, - have A : mdifferentiable_at I I.tangent (λ x, @total_space_mk M (tangent_space I) x 0) x, - { have : smooth I (I.prod 𝓘(𝕜, E)) (zero_section (tangent_space I : M → Type*)) := + have A : mdifferentiable_at I I.tangent (λ x, @total_space.mk M E (tangent_space I) x 0) x, + { have : smooth I (I.prod 𝓘(𝕜, E)) (zero_section E (tangent_space I : M → Type*)) := bundle.smooth_zero_section 𝕜 (tangent_space I : M → Type*), exact this.mdifferentiable_at }, have B : fderiv_within 𝕜 (λ (x' : E), (x', (0 : E))) (set.range ⇑I) (I ((chart_at H x) x)) v @@ -618,12 +619,12 @@ begin { exact differentiable_at_const _ }, { exact model_with_corners.unique_diff_at_image I }, { exact differentiable_at_id'.prod (differentiable_at_const _) } }, - simp only [bundle.zero_section, tangent_map, mfderiv, total_space.proj_mk, A, - if_pos, chart_at, fiber_bundle.charted_space_chart_at, tangent_bundle.trivialization_at_apply, + simp only [bundle.zero_section, tangent_map, mfderiv, A, if_pos, chart_at, + fiber_bundle.charted_space_chart_at, tangent_bundle.trivialization_at_apply, tangent_bundle_core, function.comp, continuous_linear_map.map_zero] with mfld_simps, rw [← fderiv_within_inter N] at B, rw [← fderiv_within_inter N, ← B], - congr' 2, + congr' 1, refine fderiv_within_congr (λ y hy, _) _, { simp only with mfld_simps at hy, simp only [hy, prod.mk.inj_iff] with mfld_simps }, diff --git a/src/geometry/manifold/mfderiv.lean b/src/geometry/manifold/mfderiv.lean index 30c9eb15fbf17..c621b85eee4fc 100644 --- a/src/geometry/manifold/mfderiv.lean +++ b/src/geometry/manifold/mfderiv.lean @@ -688,15 +688,9 @@ end @[simp, mfld_simps] lemma tangent_map_within_proj {p : tangent_bundle I M} : (tangent_map_within I I' f s p).proj = f p.proj := rfl -@[simp, mfld_simps] lemma tangent_map_within_fst {p : tangent_bundle I M} : - (tangent_map_within I I' f s p).1 = f p.1 := rfl - @[simp, mfld_simps] lemma tangent_map_proj {p : tangent_bundle I M} : (tangent_map I I' f p).proj = f p.proj := rfl -@[simp, mfld_simps] lemma tangent_map_fst {p : tangent_bundle I M} : - (tangent_map I I' f p).1 = f p.1 := rfl - omit Is I's lemma mdifferentiable_within_at.prod_mk {f : M → M'} {g : M → M''} @@ -841,8 +835,7 @@ lemma tangent_map_within_congr (h : ∀ x ∈ s, f x = f₁ x) (p : tangent_bundle I M) (hp : p.1 ∈ s) (hs : unique_mdiff_within_at I s p.1) : tangent_map_within I I' f s p = tangent_map_within I I' f₁ s p := begin - simp only [tangent_map_within, h p.fst hp, true_and, eq_self_iff_true, heq_iff_eq, - sigma.mk.inj_iff], + simp only [tangent_map_within, h p.1 hp, true_and, eq_self_iff_true, heq_iff_eq], congr' 1, exact mfderiv_within_congr hs h (h _ hp) end @@ -1344,17 +1337,17 @@ lemma mfderiv_within_fst {s : set (M × M')} {x : M × M'} by { rw mdifferentiable.mfderiv_within (mdifferentiable_at_fst I I') hxs, exact mfderiv_fst I I' } @[simp, mfld_simps] lemma tangent_map_prod_fst {p : tangent_bundle (I.prod I') (M × M')} : - tangent_map (I.prod I') I prod.fst p = total_space_mk p.proj.1 p.2.1 := + tangent_map (I.prod I') I prod.fst p = ⟨p.proj.1, p.2.1⟩ := by simp [tangent_map] lemma tangent_map_within_prod_fst {s : set (M × M')} {p : tangent_bundle (I.prod I') (M × M')} (hs : unique_mdiff_within_at (I.prod I') s p.proj) : - tangent_map_within (I.prod I') I prod.fst s p = total_space_mk p.proj.1 p.2.1 := + tangent_map_within (I.prod I') I prod.fst s p = ⟨p.proj.1, p.2.1⟩ := begin simp only [tangent_map_within], - rw mfderiv_within_fst, - { rcases p, refl }, - { exact hs } + rw mfderiv_within_fst _ _ hs, + rcases p, + exact ⟨rfl, heq.rfl⟩ end lemma has_mfderiv_at_snd (x : M × M') : @@ -1400,16 +1393,16 @@ lemma mfderiv_within_snd {s : set (M × M')} {x : M × M'} by { rw mdifferentiable.mfderiv_within (mdifferentiable_at_snd I I') hxs, exact mfderiv_snd I I' } @[simp, mfld_simps] lemma tangent_map_prod_snd {p : tangent_bundle (I.prod I') (M × M')} : - tangent_map (I.prod I') I' prod.snd p = total_space_mk p.proj.2 p.2.2 := + tangent_map (I.prod I') I' prod.snd p = ⟨p.proj.2, p.2.2⟩ := by simp [tangent_map] lemma tangent_map_within_prod_snd {s : set (M × M')} {p : tangent_bundle (I.prod I') (M × M')} (hs : unique_mdiff_within_at (I.prod I') s p.proj) : - tangent_map_within (I.prod I') I' prod.snd s p = total_space_mk p.proj.2 p.2.2 := + tangent_map_within (I.prod I') I' prod.snd s p = ⟨p.proj.2, p.2.2⟩ := begin simp only [tangent_map_within], rw mfderiv_within_snd, - { rcases p, refl }, + { rcases p, split; refl }, { exact hs } end @@ -1716,7 +1709,7 @@ mdifferentiable_of_mem_atlas _ (chart_mem_atlas _ _) the identification between the tangent bundle of the model space and the product space. -/ lemma tangent_map_chart {p q : tangent_bundle I M} (h : q.1 ∈ (chart_at H p.1).source) : tangent_map I I (chart_at H p.1) q = - (equiv.sigma_equiv_prod _ _).symm + (total_space.to_prod _ _).symm ((chart_at (model_prod H E) p : tangent_bundle I M → model_prod H E) q) := begin dsimp [tangent_map], @@ -1732,14 +1725,14 @@ lemma tangent_map_chart_symm {p : tangent_bundle I M} {q : tangent_bundle I H} (h : q.1 ∈ (chart_at H p.1).target) : tangent_map I I (chart_at H p.1).symm q = ((chart_at (model_prod H E) p).symm : model_prod H E → tangent_bundle I M) - ((equiv.sigma_equiv_prod H E) q) := + ((total_space.to_prod H E) q) := begin dsimp only [tangent_map], rw mdifferentiable_at.mfderiv (mdifferentiable_at_atlas_symm _ (chart_mem_atlas _ _) h), -- a trivial instance is needed after the rewrite, handle it right now. rotate, { apply_instance }, simp only [continuous_linear_map.coe_coe, tangent_bundle.chart_at, h, - tangent_bundle_core, chart_at, sigma.mk.inj_iff] with mfld_simps, + tangent_bundle_core, chart_at, total_space.to_prod_apply] with mfld_simps, end end charts @@ -2004,23 +1997,23 @@ end open bundle variables {F : Type*} [normed_add_comm_group F] [normed_space 𝕜 F] - (Z : M → Type*) [topological_space (total_space Z)] [∀ b, topological_space (Z b)] + (Z : M → Type*) [topological_space (total_space F Z)] [∀ b, topological_space (Z b)] [∀ b, add_comm_monoid (Z b)] [∀ b, module 𝕜 (Z b)] [fiber_bundle F Z] [vector_bundle 𝕜 F Z] [smooth_vector_bundle F Z I] /-- In a smooth fiber bundle, the preimage under the projection of a set with unique differential in the basis also has unique differential. -/ lemma unique_mdiff_on.smooth_bundle_preimage (hs : unique_mdiff_on I s) : - unique_mdiff_on (I.prod (𝓘(𝕜, F))) (π Z ⁻¹' s) := + unique_mdiff_on (I.prod (𝓘(𝕜, F))) (π F Z ⁻¹' s) := begin /- Using a chart (and the fact that unique differentiability is invariant under charts), we reduce the situation to the model space, where we can use the fact that products respect unique differentiability. -/ assume p hp, - replace hp : p.fst ∈ s, by simpa only with mfld_simps using hp, + replace hp : p.1 ∈ s, by simpa only with mfld_simps using hp, let e₀ := chart_at H p.1, let e := chart_at (model_prod H F) p, - have h2s : ∀ x, x ∈ e.target ∩ e.symm ⁻¹' (π Z ⁻¹' s) ↔ + have h2s : ∀ x, x ∈ e.target ∩ e.symm ⁻¹' (π F Z ⁻¹' s) ↔ (x.1 ∈ e₀.target ∧ (e₀.symm) x.1 ∈ (trivialization_at F Z p.1).base_set) ∧ (e₀.symm) x.1 ∈ s, { intro x, have A : x ∈ e.target ↔ x.1 ∈ e₀.target ∧ @@ -2032,13 +2025,13 @@ begin simp only [fiber_bundle.charted_space_chart_at_symm_fst p x hx] with mfld_simps }, -- It suffices to prove unique differentiability in a chart suffices h : unique_mdiff_on (I.prod (𝓘(𝕜, F))) - (e.target ∩ e.symm ⁻¹' (π Z ⁻¹' s)), + (e.target ∩ e.symm ⁻¹' (π F Z ⁻¹' s)), { have A : unique_mdiff_on (I.prod (𝓘(𝕜, F))) (e.symm.target ∩ - e.symm.symm ⁻¹' (e.target ∩ e.symm⁻¹' (π Z ⁻¹' s))), + e.symm.symm ⁻¹' (e.target ∩ e.symm⁻¹' (π F Z ⁻¹' s))), { apply h.unique_mdiff_on_preimage, exact (mdifferentiable_of_mem_atlas _ (chart_mem_atlas _ _)).symm, apply_instance }, - have : p ∈ e.symm.target ∩ e.symm.symm ⁻¹' (e.target ∩ e.symm⁻¹' (π Z ⁻¹' s)), + have : p ∈ e.symm.target ∩ e.symm.symm ⁻¹' (e.target ∩ e.symm⁻¹' (π F Z ⁻¹' s)), { simp only [e, hp] with mfld_simps }, apply (A _ this).mono, assume q hq, @@ -2047,7 +2040,7 @@ begin assume q hq, simp only [unique_mdiff_within_at, model_with_corners.prod, -preimage_inter] with mfld_simps, have : 𝓝[(I.symm ⁻¹' (e₀.target ∩ e₀.symm⁻¹' s) ∩ range I) ×ˢ univ] (I q.1, q.2) ≤ - 𝓝[(λ (p : E × F), (I.symm p.1, p.snd)) ⁻¹' (e.target ∩ e.symm ⁻¹' (π Z ⁻¹' s)) ∩ + 𝓝[(λ (p : E × F), (I.symm p.1, p.snd)) ⁻¹' (e.target ∩ e.symm ⁻¹' (π F Z ⁻¹' s)) ∩ (range I ×ˢ univ)] (I q.1, q.2), { rw [nhds_within_le_iff, mem_nhds_within], refine ⟨(λ (p : E × F), (I.symm p.1, p.snd)) ⁻¹' e.target, _, _, _⟩, @@ -2076,7 +2069,7 @@ begin end lemma unique_mdiff_on.tangent_bundle_proj_preimage (hs : unique_mdiff_on I s): - unique_mdiff_on I.tangent (π (tangent_space I) ⁻¹' s) := + unique_mdiff_on I.tangent (π E (tangent_space I) ⁻¹' s) := hs.smooth_bundle_preimage _ end unique_mdiff diff --git a/src/geometry/manifold/vector_bundle/basic.lean b/src/geometry/manifold/vector_bundle/basic.lean index 49445ccc558dc..2ee022460ef97 100644 --- a/src/geometry/manifold/vector_bundle/basic.lean +++ b/src/geometry/manifold/vector_bundle/basic.lean @@ -69,14 +69,14 @@ variables {𝕜 B B' F M : Type*} {E : B → Type*} /-! ### Charted space structure on a fiber bundle -/ section -variables [topological_space F] [topological_space (total_space E)] [∀ x, topological_space (E x)] +variables [topological_space F] [topological_space (total_space F E)] [∀ x, topological_space (E x)] {HB : Type*} [topological_space HB] [topological_space B] [charted_space HB B] [fiber_bundle F E] /-- A fiber bundle `E` over a base `B` with model fiber `F` is naturally a charted space modelled on `B × F`. -/ -instance fiber_bundle.charted_space : charted_space (B × F) (total_space E) := -{ atlas := (λ e : trivialization F (π E), e.to_local_homeomorph) '' trivialization_atlas F E, +instance fiber_bundle.charted_space : charted_space (B × F) (total_space F E) := +{ atlas := (λ e : trivialization F (π F E), e.to_local_homeomorph) '' trivialization_atlas F E, chart_at := λ x, (trivialization_at F E x.proj).to_local_homeomorph, mem_chart_source := λ x, (trivialization_at F E x.proj).mem_source.mpr (mem_base_set_trivialization_at F E x.proj), @@ -87,11 +87,11 @@ local attribute [reducible] model_prod /-- Let `B` be a charted space modelled on `HB`. Then a fiber bundle `E` over a base `B` with model fiber `F` is naturally a charted space modelled on `HB.prod F`. -/ -instance fiber_bundle.charted_space' : charted_space (model_prod HB F) (total_space E) := +instance fiber_bundle.charted_space' : charted_space (model_prod HB F) (total_space F E) := charted_space.comp _ (model_prod B F) _ end -lemma fiber_bundle.charted_space_chart_at (x : total_space E) : +lemma fiber_bundle.charted_space_chart_at (x : total_space F E) : chart_at (model_prod HB F) x = (trivialization_at F E x.proj).to_local_homeomorph ≫ₕ (chart_at HB x.proj).prod (local_homeomorph.refl F) := @@ -102,7 +102,7 @@ begin trivialization.coe_fst' _ (mem_base_set_trivialization_at F E x.proj)] end -lemma fiber_bundle.charted_space_chart_at_symm_fst (x : total_space E) (y : model_prod HB F) +lemma fiber_bundle.charted_space_chart_at_symm_fst (x : total_space F E) (y : model_prod HB F) (hy : y ∈ (chart_at (model_prod HB F) x).target) : ((chart_at (model_prod HB F) x).symm y).proj = (chart_at HB x.proj).symm y.1 := begin @@ -115,7 +115,7 @@ end section variables [nontrivially_normed_field 𝕜] [normed_add_comm_group F] [normed_space 𝕜 F] - [topological_space (total_space E)] [∀ x, topological_space (E x)] + [topological_space (total_space F E)] [∀ x, topological_space (E x)] {EB : Type*} [normed_add_comm_group EB] [normed_space 𝕜 EB] {HB : Type*} [topological_space HB] (IB : model_with_corners 𝕜 EB HB) @@ -127,7 +127,7 @@ variables [nontrivially_normed_field 𝕜] variables [topological_space B] [charted_space HB B] [fiber_bundle F E] -protected lemma fiber_bundle.ext_chart_at (x : total_space E) : +protected lemma fiber_bundle.ext_chart_at (x : total_space F E) : ext_chart_at (IB.prod 𝓘(𝕜, F)) x = (trivialization_at F E x.proj).to_local_equiv ≫ (ext_chart_at IB x.proj).prod (local_equiv.refl F) := @@ -146,7 +146,7 @@ namespace bundle variables {F E IB} /-- Characterization of C^n functions into a smooth vector bundle. -/ -lemma cont_mdiff_within_at_total_space (f : M → total_space E) {s : set M} {x₀ : M} : +lemma cont_mdiff_within_at_total_space (f : M → total_space F E) {s : set M} {x₀ : M} : cont_mdiff_within_at IM (IB.prod (𝓘(𝕜, F))) n f s x₀ ↔ cont_mdiff_within_at IM IB n (λ x, (f x).proj) s x₀ ∧ cont_mdiff_within_at IM 𝓘(𝕜, F) n (λ x, (trivialization_at F E (f x₀).proj (f x)).2) s x₀ := @@ -171,7 +171,7 @@ begin end /-- Characterization of C^n functions into a smooth vector bundle. -/ -lemma cont_mdiff_at_total_space (f : M → total_space E) (x₀ : M) : +lemma cont_mdiff_at_total_space (f : M → total_space F E) (x₀ : M) : cont_mdiff_at IM (IB.prod (𝓘(𝕜, F))) n f x₀ ↔ cont_mdiff_at IM IB n (λ x, (f x).proj) x₀ ∧ cont_mdiff_at IM 𝓘(𝕜, F) n (λ x, (trivialization_at F E (f x₀).proj (f x)).2) x₀ := @@ -179,12 +179,12 @@ by { simp_rw [← cont_mdiff_within_at_univ], exact cont_mdiff_within_at_total_s /-- Characterization of C^n sections of a smooth vector bundle. -/ lemma cont_mdiff_at_section (s : Π x, E x) (x₀ : B) : - cont_mdiff_at IB (IB.prod (𝓘(𝕜, F))) n (λ x, total_space_mk x (s x)) x₀ ↔ - cont_mdiff_at IB 𝓘(𝕜, F) n (λ x, (trivialization_at F E x₀ (total_space_mk x (s x))).2) x₀ := + cont_mdiff_at IB (IB.prod (𝓘(𝕜, F))) n (λ x, total_space.mk' F x (s x)) x₀ ↔ + cont_mdiff_at IB 𝓘(𝕜, F) n (λ x, (trivialization_at F E x₀ (total_space.mk' F x (s x))).2) x₀ := by { simp_rw [cont_mdiff_at_total_space, and_iff_right_iff_imp], intro x, exact cont_mdiff_at_id } variables (E) -lemma cont_mdiff_proj : cont_mdiff (IB.prod 𝓘(𝕜, F)) IB n (π E) := +lemma cont_mdiff_proj : cont_mdiff (IB.prod 𝓘(𝕜, F)) IB n (π F E) := begin intro x, rw [cont_mdiff_at, cont_mdiff_within_at_iff'], @@ -193,47 +193,47 @@ begin apply cont_diff_within_at_fst.congr, { rintros ⟨a, b⟩ hab, simp only with mfld_simps at hab, - have : ((chart_at HB x.1).symm (IB.symm a), b) ∈ (trivialization_at F E x.fst).target, + have : ((chart_at HB x.1).symm (IB.symm a), b) ∈ (trivialization_at F E x.proj).target, { simp only [hab] with mfld_simps }, simp only [trivialization.proj_symm_apply _ this, hab] with mfld_simps }, { simp only with mfld_simps } end -lemma smooth_proj : smooth (IB.prod 𝓘(𝕜, F)) IB (π E) := +lemma smooth_proj : smooth (IB.prod 𝓘(𝕜, F)) IB (π F E) := cont_mdiff_proj E -lemma cont_mdiff_on_proj {s : set (total_space E)} : - cont_mdiff_on (IB.prod 𝓘(𝕜, F)) IB n (π E) s := +lemma cont_mdiff_on_proj {s : set (total_space F E)} : + cont_mdiff_on (IB.prod 𝓘(𝕜, F)) IB n (π F E) s := (bundle.cont_mdiff_proj E).cont_mdiff_on -lemma smooth_on_proj {s : set (total_space E)} : - smooth_on (IB.prod 𝓘(𝕜, F)) IB (π E) s := +lemma smooth_on_proj {s : set (total_space F E)} : + smooth_on (IB.prod 𝓘(𝕜, F)) IB (π F E) s := cont_mdiff_on_proj E -lemma cont_mdiff_at_proj {p : total_space E} : +lemma cont_mdiff_at_proj {p : total_space F E} : cont_mdiff_at (IB.prod 𝓘(𝕜, F)) IB n - (π E) p := + (π F E) p := (bundle.cont_mdiff_proj E).cont_mdiff_at -lemma smooth_at_proj {p : total_space E} : - smooth_at (IB.prod 𝓘(𝕜, F)) IB (π E) p := +lemma smooth_at_proj {p : total_space F E} : + smooth_at (IB.prod 𝓘(𝕜, F)) IB (π F E) p := bundle.cont_mdiff_at_proj E lemma cont_mdiff_within_at_proj - {s : set (total_space E)} - {p : total_space E} : - cont_mdiff_within_at (IB.prod 𝓘(𝕜, F)) IB n (π E) s p := + {s : set (total_space F E)} + {p : total_space F E} : + cont_mdiff_within_at (IB.prod 𝓘(𝕜, F)) IB n (π F E) s p := (bundle.cont_mdiff_at_proj E).cont_mdiff_within_at lemma smooth_within_at_proj - {s : set (total_space E)} - {p : total_space E} : - smooth_within_at (IB.prod 𝓘(𝕜, F)) IB (π E) s p := + {s : set (total_space F E)} + {p : total_space F E} : + smooth_within_at (IB.prod 𝓘(𝕜, F)) IB (π F E) s p := bundle.cont_mdiff_within_at_proj E variables (𝕜 E) [∀ x, add_comm_monoid (E x)] [∀ x, module 𝕜 (E x)] [vector_bundle 𝕜 F E] -lemma smooth_zero_section : smooth IB (IB.prod 𝓘(𝕜, F)) (zero_section E) := +lemma smooth_zero_section : smooth IB (IB.prod 𝓘(𝕜, F)) (zero_section F E) := begin intro x, rw [bundle.cont_mdiff_at_total_space], @@ -263,7 +263,7 @@ variables [nontrivially_normed_field 𝕜] section with_topology -variables [topological_space (total_space E)] [∀ x, topological_space (E x)] +variables [topological_space (total_space F E)] [∀ x, topological_space (E x)] variables (F E) [fiber_bundle F E] [vector_bundle 𝕜 F E] @@ -272,7 +272,7 @@ topological vector bundle over `B` with fibers isomorphic to `F`, then `smooth_v registers that the bundle is smooth, in the sense of having smooth transition functions. This is a mixin, not carrying any new data`. -/ class smooth_vector_bundle : Prop := -(smooth_on_coord_change : ∀ (e e' : trivialization F (π E)) +(smooth_on_coord_change : ∀ (e e' : trivialization F (π F E)) [mem_trivialization_atlas e] [mem_trivialization_atlas e'], smooth_on IB 𝓘(𝕜, F →L[𝕜] F) (λ b : B, (e.coord_changeL 𝕜 e' b : F →L[𝕜] F)) (e.base_set ∩ e'.base_set)) @@ -284,7 +284,7 @@ variables [smooth_vector_bundle F E IB] /-- For a smooth vector bundle `E` over `B` with fiber modelled on `F`, the change-of-co-ordinates between two trivializations `e`, `e'` for `E`, considered as charts to `B × F`, is smooth and fiberwise linear. -/ -instance : has_groupoid (total_space E) (smooth_fiberwise_linear B F IB) := +instance : has_groupoid (total_space F E) (smooth_fiberwise_linear B F IB) := { compatible := begin rintros _ _ ⟨e, he, rfl⟩ ⟨e', he', rfl⟩, haveI : mem_trivialization_atlas e := ⟨he⟩, @@ -308,7 +308,7 @@ instance : has_groupoid (total_space E) (smooth_fiberwise_linear B F IB) := end } /-- A smooth vector bundle `E` is naturally a smooth manifold. -/ -instance : smooth_manifold_with_corners (IB.prod 𝓘(𝕜, F)) (total_space E) := +instance : smooth_manifold_with_corners (IB.prod 𝓘(𝕜, F)) (total_space F E) := begin refine { .. structure_groupoid.has_groupoid.comp (smooth_fiberwise_linear B F IB) _ }, intros e he, @@ -366,11 +366,11 @@ instance bundle.trivial.smooth_vector_bundle : smooth_vector_bundle F (bundle.tr section prod variables (F₁ : Type*) [normed_add_comm_group F₁] [normed_space 𝕜 F₁] - (E₁ : B → Type*) [topological_space (total_space E₁)] + (E₁ : B → Type*) [topological_space (total_space F₁ E₁)] [Π x, add_comm_monoid (E₁ x)] [Π x, module 𝕜 (E₁ x)] variables (F₂ : Type*) [normed_add_comm_group F₂] [normed_space 𝕜 F₂] - (E₂ : B → Type*) [topological_space (total_space E₂)] + (E₂ : B → Type*) [topological_space (total_space F₂ E₂)] [Π x, add_comm_monoid (E₂ x)] [Π x, module 𝕜 (E₂ x)] variables [Π x : B, topological_space (E₁ x)] [Π x : B, topological_space (E₂ x)] @@ -409,10 +409,9 @@ variables [∀ x, topological_space (E x)] {F E} class is_smooth (a : vector_prebundle 𝕜 F E) : Prop := (exists_smooth_coord_change : ∀ (e e' ∈ a.pretrivialization_atlas), ∃ f : B → F →L[𝕜] F, smooth_on IB 𝓘(𝕜, F →L[𝕜] F) f (e.base_set ∩ e'.base_set) ∧ - ∀ (b : B) (hb : b ∈ e.base_set ∩ e'.base_set) (v : F), - f b v = (e' (total_space_mk b (e.symm b v))).2) + ∀ (b : B) (hb : b ∈ e.base_set ∩ e'.base_set) (v : F), f b v = (e' ⟨b ,e.symm b v⟩).2) -variables (a : vector_prebundle 𝕜 F E) [ha : a.is_smooth IB] {e e' : pretrivialization F (π E)} +variables (a : vector_prebundle 𝕜 F E) [ha : a.is_smooth IB] {e e' : pretrivialization F (π F E)} include ha /-- A randomly chosen coordinate change on a `smooth_vector_prebundle`, given by @@ -430,12 +429,12 @@ lemma smooth_on_smooth_coord_change (he : e ∈ a.pretrivialization_atlas) lemma smooth_coord_change_apply (he : e ∈ a.pretrivialization_atlas) (he' : e' ∈ a.pretrivialization_atlas) {b : B} (hb : b ∈ e.base_set ∩ e'.base_set) (v : F) : - a.smooth_coord_change IB he he' b v = (e' (total_space_mk b (e.symm b v))).2 := + a.smooth_coord_change IB he he' b v = (e' ⟨b, e.symm b v⟩).2 := (classical.some_spec (ha.exists_smooth_coord_change e he e' he')).2 b hb v lemma mk_smooth_coord_change (he : e ∈ a.pretrivialization_atlas) (he' : e' ∈ a.pretrivialization_atlas) {b : B} (hb : b ∈ e.base_set ∩ e'.base_set) (v : F) : - (b, (a.smooth_coord_change IB he he' b v)) = e' (total_space_mk b (e.symm b v)) := + (b, (a.smooth_coord_change IB he he' b v)) = e' ⟨b, e.symm b v⟩ := begin ext, { rw [e.mk_symm hb.1 v, e'.coe_fst', e.proj_symm_apply' hb.1], diff --git a/src/geometry/manifold/vector_bundle/hom.lean b/src/geometry/manifold/vector_bundle/hom.lean index e0930af6bb21f..e83f070f9d9a6 100644 --- a/src/geometry/manifold/vector_bundle/hom.lean +++ b/src/geometry/manifold/vector_bundle/hom.lean @@ -25,13 +25,13 @@ variables {𝕜 B F F₁ F₂ M M₁ M₂ : Type*} [nontrivially_normed_field 𝕜] [∀ x, add_comm_group (E x)] [∀ x, module 𝕜 (E x)] [normed_add_comm_group F] [normed_space 𝕜 F] - [topological_space (total_space E)] [∀ x, topological_space (E x)] + [topological_space (total_space F E)] [∀ x, topological_space (E x)] [∀ x, add_comm_group (E₁ x)] [∀ x, module 𝕜 (E₁ x)] [normed_add_comm_group F₁] [normed_space 𝕜 F₁] - [topological_space (total_space E₁)] [∀ x, topological_space (E₁ x)] + [topological_space (total_space F₁ E₁)] [∀ x, topological_space (E₁ x)] [∀ x, add_comm_group (E₂ x)] [∀ x, module 𝕜 (E₂ x)] [normed_add_comm_group F₂] [normed_space 𝕜 F₂] - [topological_space (total_space E₂)] [∀ x, topological_space (E₂ x)] + [topological_space (total_space F₂ E₂)] [∀ x, topological_space (E₂ x)] [_i₁ : ∀ x, topological_add_group (E₂ x)] [_i₂ : ∀ x, has_continuous_smul 𝕜 (E₂ x)] {EB : Type*} [normed_add_comm_group EB] [normed_space 𝕜 EB] @@ -43,9 +43,10 @@ variables {𝕜 B F F₁ F₂ M M₁ M₂ : Type*} {n : ℕ∞} [fiber_bundle F₁ E₁] [vector_bundle 𝕜 F₁ E₁] [fiber_bundle F₂ E₂] [vector_bundle 𝕜 F₂ E₂] - {e₁ e₁' : trivialization F₁ (π E₁)} {e₂ e₂' : trivialization F₂ (π E₂)} + {e₁ e₁' : trivialization F₁ (π F₁ E₁)} {e₂ e₂' : trivialization F₂ (π F₂ E₂)} -local notation `LE₁E₂` := total_space (bundle.continuous_linear_map (ring_hom.id 𝕜) F₁ E₁ F₂ E₂) +local notation `LE₁E₂` := total_space (F₁ →L[𝕜] F₂) + (bundle.continuous_linear_map (ring_hom.id 𝕜) E₁ E₂) /- This proof is slow, especially the `simp only` and the elaboration of `h₂`. -/ lemma smooth_on_continuous_linear_map_coord_change @@ -110,6 +111,5 @@ instance bundle.continuous_linear_map.vector_prebundle.is_smooth : end } instance smooth_vector_bundle.continuous_linear_map : - smooth_vector_bundle (F₁ →L[𝕜] F₂) (bundle.continuous_linear_map (ring_hom.id 𝕜) F₁ E₁ F₂ E₂) - IB := + smooth_vector_bundle (F₁ →L[𝕜] F₂) (bundle.continuous_linear_map (ring_hom.id 𝕜) E₁ E₂) IB := (bundle.continuous_linear_map.vector_prebundle (ring_hom.id 𝕜) F₁ E₁ F₂ E₂).smooth_vector_bundle IB diff --git a/src/geometry/manifold/vector_bundle/pullback.lean b/src/geometry/manifold/vector_bundle/pullback.lean index 44c3c92709f7c..a68fa9c52a52a 100644 --- a/src/geometry/manifold/vector_bundle/pullback.lean +++ b/src/geometry/manifold/vector_bundle/pullback.lean @@ -27,7 +27,7 @@ variables {𝕜 B B' M : Type*} (F : Type*) (E : B → Type*) variables [nontrivially_normed_field 𝕜] [∀ x, add_comm_monoid (E x)] [∀ x, module 𝕜 (E x)] [normed_add_comm_group F] [normed_space 𝕜 F] - [topological_space (total_space E)] [∀ x, topological_space (E x)] + [topological_space (total_space F E)] [∀ x, topological_space (E x)] {EB : Type*} [normed_add_comm_group EB] [normed_space 𝕜 EB] {HB : Type*} [topological_space HB] (IB : model_with_corners 𝕜 EB HB) diff --git a/src/geometry/manifold/vector_bundle/smooth_section.lean b/src/geometry/manifold/vector_bundle/smooth_section.lean index 9190ea480622e..c7782d18cc4b6 100644 --- a/src/geometry/manifold/vector_bundle/smooth_section.lean +++ b/src/geometry/manifold/vector_bundle/smooth_section.lean @@ -15,7 +15,7 @@ In this file we define the type `cont_mdiff_section` of `n` times continuously d sections of a smooth vector bundle over a manifold `M` and prove that it's a module. -/ open bundle filter function -open_locale manifold +open_locale bundle manifold variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] {E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] @@ -34,7 +34,7 @@ variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] variables (F : Type*) [normed_add_comm_group F] [normed_space 𝕜 F] -- `F` model fiber (n : ℕ∞) - (V : M → Type*) [topological_space (total_space V)] -- `V` vector bundle + (V : M → Type*) [topological_space (total_space F V)] -- `V` vector bundle [Π x, add_comm_group (V x)] [Π x, module 𝕜 (V x)] variables [Π x : M, topological_space (V x)] [fiber_bundle F V] @@ -45,7 +45,7 @@ variables [Π x : M, topological_space (V x)] @[protect_proj] structure cont_mdiff_section := (to_fun : Π x, V x) -(cont_mdiff_to_fun : cont_mdiff I (I.prod 𝓘(𝕜, F)) n (λ x, total_space_mk x (to_fun x))) +(cont_mdiff_to_fun : cont_mdiff I (I.prod 𝓘(𝕜, F)) n (λ x, total_space.mk' F x (to_fun x))) /-- Bundled smooth sections of a vector bundle. -/ @[reducible] def smooth_section := cont_mdiff_section I F ⊤ V @@ -62,26 +62,26 @@ instance : has_coe_to_fun Cₛ^n⟮I; F, V⟯ (λ s, Π x, V x) := ⟨cont_mdiff variables {s t : Cₛ^n⟮I; F, V⟯} @[simp] lemma coe_fn_mk (s : Π x, V x) - (hs : cont_mdiff I (I.prod 𝓘(𝕜, F)) n (λ x, total_space_mk x (s x))) : + (hs : cont_mdiff I (I.prod 𝓘(𝕜, F)) n (λ x, total_space.mk x (s x))) : (mk s hs : Π x, V x) = s := rfl protected lemma cont_mdiff (s : Cₛ^n⟮I; F, V⟯) : - cont_mdiff I (I.prod 𝓘(𝕜, F)) n (λ x, total_space_mk x (s x : V x)) := s.cont_mdiff_to_fun + cont_mdiff I (I.prod 𝓘(𝕜, F)) n (λ x, total_space.mk' F x (s x : V x)) := s.cont_mdiff_to_fun protected lemma smooth (s : Cₛ^∞⟮I; F, V⟯) : - smooth I (I.prod 𝓘(𝕜, F)) (λ x, total_space_mk x (s x : V x)) := s.cont_mdiff_to_fun + smooth I (I.prod 𝓘(𝕜, F)) (λ x, total_space.mk' F x (s x : V x)) := s.cont_mdiff_to_fun protected lemma mdifferentiable' (s : Cₛ^n⟮I; F, V⟯) (hn : 1 ≤ n) : - mdifferentiable I (I.prod 𝓘(𝕜, F)) (λ x, total_space_mk x (s x : V x)) := + mdifferentiable I (I.prod 𝓘(𝕜, F)) (λ x, total_space.mk' F x (s x : V x)) := s.cont_mdiff.mdifferentiable hn protected lemma mdifferentiable (s : Cₛ^∞⟮I; F, V⟯) : - mdifferentiable I (I.prod 𝓘(𝕜, F)) (λ x, total_space_mk x (s x : V x)) := + mdifferentiable I (I.prod 𝓘(𝕜, F)) (λ x, total_space.mk' F x (s x : V x)) := s.cont_mdiff.mdifferentiable le_top protected lemma mdifferentiable_at (s : Cₛ^∞⟮I; F, V⟯) {x} : - mdifferentiable_at I (I.prod 𝓘(𝕜, F)) (λ x, total_space_mk x (s x : V x)) x := + mdifferentiable_at I (I.prod 𝓘(𝕜, F)) (λ x, total_space.mk' F x (s x : V x)) x := s.mdifferentiable x lemma coe_inj ⦃s t : Cₛ^n⟮I; F, V⟯⦄ (h : (s : Π x, V x) = t) : s = t := diff --git a/src/geometry/manifold/vector_bundle/tangent.lean b/src/geometry/manifold/vector_bundle/tangent.lean index 55d597435cb96..491f7107ff786 100644 --- a/src/geometry/manifold/vector_bundle/tangent.lean +++ b/src/geometry/manifold/vector_bundle/tangent.lean @@ -136,11 +136,10 @@ variable (M) /-- The tangent bundle to a smooth manifold, as a Sigma type. Defined in terms of `bundle.total_space` to be able to put a suitable topology on it. -/ @[nolint has_nonempty_instance, reducible] -- is empty if the base manifold is empty -def tangent_bundle := bundle.total_space (tangent_space I : M → Type*) +def tangent_bundle := bundle.total_space E (tangent_space I : M → Type*) local notation `TM` := tangent_bundle I M - section tangent_bundle_instances /- In general, the definition of tangent_space is not reducible, so that type class inference @@ -187,7 +186,8 @@ rfl @[simp, mfld_simps] lemma trivialization_at_source (x : M) : - (trivialization_at E (tangent_space I) x).source = π _ ⁻¹' (chart_at H x).source := + (trivialization_at E (tangent_space I) x).source = + π E (tangent_space I) ⁻¹' (chart_at H x).source := rfl @[simp, mfld_simps] @@ -287,7 +287,7 @@ end tangent_bundle_instances /-- In the tangent bundle to the model space, the charts are just the canonical identification between a product type and a sigma type, a.k.a. `equiv.sigma_equiv_prod`. -/ @[simp, mfld_simps] lemma tangent_bundle_model_space_chart_at (p : tangent_bundle I H) : - (chart_at (model_prod H E) p).to_local_equiv = (equiv.sigma_equiv_prod H E).to_local_equiv := + (chart_at (model_prod H E) p).to_local_equiv = (total_space.to_prod H E).to_local_equiv := begin ext x : 1, { ext, { refl }, @@ -303,12 +303,12 @@ begin end @[simp, mfld_simps] lemma tangent_bundle_model_space_coe_chart_at (p : tangent_bundle I H) : - ⇑(chart_at (model_prod H E) p) = equiv.sigma_equiv_prod H E := + ⇑(chart_at (model_prod H E) p) = total_space.to_prod H E := by { unfold_coes, simp_rw [tangent_bundle_model_space_chart_at], refl } @[simp, mfld_simps] lemma tangent_bundle_model_space_coe_chart_at_symm (p : tangent_bundle I H) : ((chart_at (model_prod H E) p).symm : model_prod H E → tangent_bundle I H) = - (equiv.sigma_equiv_prod H E).symm := + (total_space.to_prod H E).symm := by { unfold_coes, simp_rw [local_homeomorph.symm_to_local_equiv, tangent_bundle_model_space_chart_at], refl } @@ -339,16 +339,16 @@ def tangent_bundle_model_space_homeomorph : tangent_bundle I H ≃ₜ model_prod simp only with mfld_simps }, simpa only with mfld_simps using this, end, - .. equiv.sigma_equiv_prod H E } + .. total_space.to_prod H E } @[simp, mfld_simps] lemma tangent_bundle_model_space_homeomorph_coe : (tangent_bundle_model_space_homeomorph H I : tangent_bundle I H → model_prod H E) - = equiv.sigma_equiv_prod H E := + = total_space.to_prod H E := rfl @[simp, mfld_simps] lemma tangent_bundle_model_space_homeomorph_coe_symm : ((tangent_bundle_model_space_homeomorph H I).symm : model_prod H E → tangent_bundle I H) - = (equiv.sigma_equiv_prod H E).symm := + = (total_space.to_prod H E).symm := rfl section in_tangent_coordinates diff --git a/src/topology/covering.lean b/src/topology/covering.lean index 58f758a2ee4ed..7abc9ef9f9938 100644 --- a/src/topology/covering.lean +++ b/src/topology/covering.lean @@ -161,7 +161,7 @@ protected lemma is_fiber_bundle.is_covering_map {F : Type*} [topological_space F is_covering_map.mk f (λ x, F) (λ x, classical.some (hf x)) (λ x, classical.some_spec (hf x)) protected lemma fiber_bundle.is_covering_map {F : Type*} {E : X → Type*} [topological_space F] - [discrete_topology F] [topological_space (bundle.total_space E)] [Π x, topological_space (E x)] - [hf : fiber_bundle F E] : is_covering_map (π E) := + [discrete_topology F] [topological_space (bundle.total_space F E)] [Π x, topological_space (E x)] + [hf : fiber_bundle F E] : is_covering_map (π F E) := is_fiber_bundle.is_covering_map (λ x, ⟨trivialization_at F E x, mem_base_set_trivialization_at F E x ⟩) diff --git a/src/topology/fiber_bundle/basic.lean b/src/topology/fiber_bundle/basic.lean index f50b7b4f4e649..e855d01461f19 100644 --- a/src/topology/fiber_bundle/basic.lean +++ b/src/topology/fiber_bundle/basic.lean @@ -16,16 +16,14 @@ Mathematically, a (topological) fiber bundle with fiber `F` over a base `B` is a point is a direct product. In our formalism, a fiber bundle is by definition the type -`bundle.total_space E` where `E : B → Type*` is a function associating to -`x : B` the fiber over `x`. This type `bundle.total_space E` is just a type synonym for -`Σ (x : B), E x`, with the interest that one can put another topology than on `Σ (x : B), E x` -which has the disjoint union topology. +`bundle.total_space F E` where `E : B → Type*` is a function associating to `x : B` the fiber over +`x`. This type `bundle.total_space F E` is a type of pairs `(proj : B, snd : E proj)`. -To have a fiber bundle structure on `bundle.total_space E`, one should +To have a fiber bundle structure on `bundle.total_space F E`, one should additionally have the following data: * `F` should be a topological space; -* There should be a topology on `bundle.total_space E`, for which the projection to `B` is +* There should be a topology on `bundle.total_space F E`, for which the projection to `B` is a fiber bundle with fiber `F` (in particular, each fiber `E x` is homeomorphic to `F`); * For each `x`, the fiber `E x` should be a topological space, and the injection from `E x` to `bundle.total_space F E` should be an embedding; @@ -52,17 +50,16 @@ fiber bundle from trivializations given as local equivalences with minimum addit ### Construction of a bundle from trivializations -* `bundle.total_space E` is a type synonym for `Σ (x : B), E x`, that we can endow with a suitable - topology. +* `bundle.total_space F E` is the type of pairs `(proj : B, snd : E proj)`. We can use the extra + argument `F` to construct topology on the total space. * `fiber_bundle_core ι B F` : structure registering how changes of coordinates act on the fiber `F` above open subsets of `B`, where local trivializations are indexed by `ι`. Let `Z : fiber_bundle_core ι B F`. Then we define * `Z.fiber x` : the fiber above `x`, homeomorphic to `F` (and defeq to `F` as a type). -* `Z.total_space` : the total space of `Z`, defined as a `Type` as `Σ (b : B), F`, but with a - twisted topology coming from the fiber bundle structure. It is (reducibly) the same as - `bundle.total_space Z.fiber`. +* `Z.total_space` : the total space of `Z`, defined as a `Type*` as `bundle.total_space F Z.fiber` + with a custom topology. * `Z.proj` : projection from `Z.total_space` to `B`. It is continuous. * `Z.local_triv i`: for `i : ι`, bundle trivialization above the set `Z.base_set i`, which is an open set in `B`. @@ -153,8 +150,8 @@ choose for each `x` one specific trivialization around it. We include this choic of the `fiber_bundle_core`, as it makes some constructions more functorial and it is a nice way to say that the trivializations cover the whole space `B`. -With this definition, the type of the fiber bundle space constructed from the core data is just -`Σ (b : B), F `, but the topology is not the product one, in general. +With this definition, the type of the fiber bundle space constructed from the core data is +`bundle.total_space F (λ b : B, F)`, but the topology is not the product one, in general. We also take the indexing type (indexing all the trivializations) as a parameter to the fiber bundle core: it could always be taken as a subtype of all the maps from open subsets of `B` to continuous @@ -172,7 +169,7 @@ variables {ι B F X : Type*} [topological_space X] open topological_space filter set bundle open_locale topology classical bundle -attribute [mfld_simps] total_space_mk coe_fst coe_snd coe_snd_map_apply +attribute [mfld_simps] total_space.coe_proj total_space.coe_snd coe_snd_map_apply coe_snd_map_smul total_space.mk_cast /-! ### General definition of fiber bundles -/ @@ -180,15 +177,15 @@ attribute [mfld_simps] total_space_mk coe_fst coe_snd coe_snd_map_apply section fiber_bundle variables (F) [topological_space B] [topological_space F] (E : B → Type*) - [topological_space (total_space E)] [∀ b, topological_space (E b)] + [topological_space (total_space F E)] [∀ b, topological_space (E b)] /-- A (topological) fiber bundle with fiber `F` over a base `B` is a space projecting on `B` for which the fibers are all homeomorphic to `F`, such that the local situation around each point is a direct product. -/ class fiber_bundle := -(total_space_mk_inducing [] : ∀ (b : B), inducing (@total_space_mk B E b)) -(trivialization_atlas [] : set (trivialization F (π E))) -(trivialization_at [] : B → trivialization F (π E)) +(total_space_mk_inducing [] : ∀ (b : B), inducing (@total_space.mk B F E b)) +(trivialization_atlas [] : set (trivialization F (π F E))) +(trivialization_at [] : B → trivialization F (π F E)) (mem_base_set_trivialization_at [] : ∀ b : B, b ∈ (trivialization_at b).base_set) (trivialization_mem_atlas [] : ∀ b : B, trivialization_at b ∈ trivialization_atlas) @@ -202,7 +199,7 @@ bundle. This is needed because lemmas about the linearity of trivializations or functions to `F →L[R] F`, where `F` is the model fiber) of the transition functions are only expected to hold for trivializations in the designated atlas. -/ @[mk_iff] -class mem_trivialization_atlas [fiber_bundle F E] (e : trivialization F (π E)) : Prop := +class mem_trivialization_atlas [fiber_bundle F E] (e : trivialization F (π F E)) : Prop := (out : e ∈ trivialization_atlas F E) instance [fiber_bundle F E] (b : B) : mem_trivialization_atlas (trivialization_at F E b) := @@ -211,44 +208,44 @@ instance [fiber_bundle F E] (b : B) : mem_trivialization_atlas (trivialization_a namespace fiber_bundle variables (F) {E} [fiber_bundle F E] -lemma map_proj_nhds (x : total_space E) : - map (π E) (𝓝 x) = 𝓝 x.proj := +lemma map_proj_nhds (x : total_space F E) : + map (π F E) (𝓝 x) = 𝓝 x.proj := (trivialization_at F E x.proj).map_proj_nhds $ (trivialization_at F E x.proj).mem_source.2 $ mem_base_set_trivialization_at F E x.proj variables (E) /-- The projection from a fiber bundle to its base is continuous. -/ -@[continuity] lemma continuous_proj : continuous (π E) := +@[continuity] lemma continuous_proj : continuous (π F E) := continuous_iff_continuous_at.2 $ λ x, (map_proj_nhds F x).le /-- The projection from a fiber bundle to its base is an open map. -/ -lemma is_open_map_proj : is_open_map (π E) := +lemma is_open_map_proj : is_open_map (π F E) := is_open_map.of_nhds_le $ λ x, (map_proj_nhds F x).ge /-- The projection from a fiber bundle with a nonempty fiber to its base is a surjective map. -/ -lemma surjective_proj [nonempty F] : function.surjective (π E) := +lemma surjective_proj [nonempty F] : function.surjective (π F E) := λ b, let ⟨p, _, hpb⟩ := (trivialization_at F E b).proj_surj_on_base_set (mem_base_set_trivialization_at F E b) in ⟨p, hpb⟩ /-- The projection from a fiber bundle with a nonempty fiber to its base is a quotient map. -/ -lemma quotient_map_proj [nonempty F] : quotient_map (π E) := +lemma quotient_map_proj [nonempty F] : quotient_map (π F E) := (is_open_map_proj F E).to_quotient_map (continuous_proj F E) (surjective_proj F E) -lemma continuous_total_space_mk (x : B) : continuous (@total_space_mk B E x) := +lemma continuous_total_space_mk (x : B) : continuous (@total_space.mk B F E x) := (total_space_mk_inducing F E x).continuous variables {E F} @[simp, mfld_simps] -lemma mem_trivialization_at_proj_source {x : total_space E} : +lemma mem_trivialization_at_proj_source {x : total_space F E} : x ∈ (trivialization_at F E x.proj).source := (trivialization.mem_source _).mpr $ mem_base_set_trivialization_at F E x.proj @[simp, mfld_simps] -lemma trivialization_at_proj_fst {x : total_space E} : +lemma trivialization_at_proj_fst {x : total_space F E} : ((trivialization_at F E x.proj) x).1 = x.proj := trivialization.coe_fst' _ $ mem_base_set_trivialization_at F E x.proj @@ -256,7 +253,7 @@ variable (F) open trivialization /-- Characterization of continuous functions (at a point, within a set) into a fiber bundle. -/ -lemma continuous_within_at_total_space (f : X → total_space E) {s : set X} {x₀ : X} : +lemma continuous_within_at_total_space (f : X → total_space F E) {s : set X} {x₀ : X} : continuous_within_at f s x₀ ↔ continuous_within_at (λ x, (f x).proj) s x₀ ∧ continuous_within_at (λ x, ((trivialization_at F E (f x₀).proj) (f x)).2) s x₀ := @@ -276,7 +273,7 @@ begin end /-- Characterization of continuous functions (at a point) into a fiber bundle. -/ -lemma continuous_at_total_space (f : X → total_space E) {x₀ : X} : +lemma continuous_at_total_space (f : X → total_space F E) {x₀ : X} : continuous_at f x₀ ↔ continuous_at (λ x, (f x).proj) x₀ ∧ continuous_at (λ x, ((trivialization_at F E (f x₀).proj) (f x)).2) x₀ := by { simp_rw [← continuous_within_at_univ], exact continuous_within_at_total_space F f } @@ -289,16 +286,16 @@ variables (F E) then it is trivial over any closed interval. -/ lemma fiber_bundle.exists_trivialization_Icc_subset [conditionally_complete_linear_order B] [order_topology B] [fiber_bundle F E] (a b : B) : - ∃ e : trivialization F (π E), Icc a b ⊆ e.base_set := + ∃ e : trivialization F (π F E), Icc a b ⊆ e.base_set := begin classical, - obtain ⟨ea, hea⟩ : ∃ ea : trivialization F (π E), a ∈ ea.base_set := + obtain ⟨ea, hea⟩ : ∃ ea : trivialization F (π F E), a ∈ ea.base_set := ⟨trivialization_at F E a, mem_base_set_trivialization_at F E a⟩, -- If `a < b`, then `[a, b] = ∅`, and the statement is trivial cases le_or_lt a b with hab hab; [skip, exact ⟨ea, by simp *⟩], /- Let `s` be the set of points `x ∈ [a, b]` such that `E` is trivializable over `[a, x]`. We need to show that `b ∈ s`. Let `c = Sup s`. We will show that `c ∈ s` and `c = b`. -/ - set s : set B := {x ∈ Icc a b | ∃ e : trivialization F (π E), Icc a x ⊆ e.base_set}, + set s : set B := {x ∈ Icc a b | ∃ e : trivialization F (π F E), Icc a x ⊆ e.base_set}, have ha : a ∈ s, from ⟨left_mem_Icc.2 hab, ea, by simp [hea]⟩, have sne : s.nonempty := ⟨a, ha⟩, have hsb : b ∈ upper_bounds s, from λ x hx, hx.1.2, @@ -306,12 +303,12 @@ begin set c := Sup s, have hsc : is_lub s c, from is_lub_cSup sne sbd, have hc : c ∈ Icc a b, from ⟨hsc.1 ha, hsc.2 hsb⟩, - obtain ⟨-, ec : trivialization F (π E), hec : Icc a c ⊆ ec.base_set⟩ : c ∈ s, + obtain ⟨-, ec : trivialization F (π F E), hec : Icc a c ⊆ ec.base_set⟩ : c ∈ s, { cases hc.1.eq_or_lt with heq hlt, { rwa ← heq }, refine ⟨hc, _⟩, /- In order to show that `c ∈ s`, consider a trivialization `ec` of `proj` over a neighborhood of `c`. Its base set includes `(c', c]` for some `c' ∈ [a, c)`. -/ - obtain ⟨ec, hc⟩ : ∃ ec : trivialization F (π E), c ∈ ec.base_set := + obtain ⟨ec, hc⟩ : ∃ ec : trivialization F (π F E), c ∈ ec.base_set := ⟨trivialization_at F E c, mem_base_set_trivialization_at F E c⟩, obtain ⟨c', hc', hc'e⟩ : ∃ c' ∈ Ico a c, Ioc c' c ⊆ ec.base_set := (mem_nhds_within_Iic_iff_exists_mem_Ico_Ioc_subset hlt).1 @@ -325,7 +322,7 @@ begin done. Otherwise we show that `proj` can be trivialized over a larger interval `[a, d]`, `d ∈ (c, b]`, hence `c` is not an upper bound of `s`. -/ cases hc.2.eq_or_lt with heq hlt, { exact ⟨ec, heq ▸ hec⟩ }, - rsuffices ⟨d, hdcb, hd⟩ : ∃ (d ∈ Ioc c b) (e : trivialization F (π E)), Icc a d ⊆ e.base_set, + rsuffices ⟨d, hdcb, hd⟩ : ∃ (d ∈ Ioc c b) (e : trivialization F (π F E)), Icc a d ⊆ e.base_set, { exact ((hsc.1 ⟨⟨hc.1.trans hdcb.1.le, hdcb.2⟩, hd⟩).not_lt hdcb.1).elim }, /- Since the base set of `ec` is open, it includes `[c, d)` (hence, `[a, d)`) for some `d ∈ (c, b]`. -/ @@ -338,7 +335,7 @@ begin { /- If `(c, d) = ∅`, then let `ed` be a trivialization of `proj` over a neighborhood of `d`. Then the disjoint union of `ec` restricted to `(-∞, d)` and `ed` restricted to `(c, ∞)` is a trivialization over `[a, d]`. -/ - obtain ⟨ed, hed⟩ : ∃ ed : trivialization F (π E), d ∈ ed.base_set := + obtain ⟨ed, hed⟩ : ∃ ed : trivialization F (π F E), d ∈ ed.base_set := ⟨trivialization_at F E d, mem_base_set_trivialization_at F E d⟩, refine ⟨d, hdcb, (ec.restr_open (Iio d) is_open_Iio).disjoint_union (ed.restr_open (Ioi c) is_open_Ioi) (he.mono (inter_subset_right _ _) @@ -396,18 +393,13 @@ typeclass inference -/ @[nolint unused_arguments has_nonempty_instance] def fiber (x : B) := F -section fiber_instances -local attribute [reducible] fiber - -instance topological_space_fiber (x : B) : topological_space (Z.fiber x) := by apply_instance - -end fiber_instances +instance topological_space_fiber (x : B) : topological_space (Z.fiber x) := +‹topological_space F› /-- The total space of the fiber bundle, as a convenience function for dot notation. -It is by definition equal to `bundle.total_space Z.fiber`, a.k.a. `Σ x, Z.fiber x` but with a -different name for typeclass inference. -/ +It is by definition equal to `bundle.total_space Z.fiber` -/ @[nolint unused_arguments, reducible] -def total_space := bundle.total_space Z.fiber +def total_space := bundle.total_space F Z.fiber /-- The projection from the total space of a fiber bundle core, on its base. -/ @[reducible, simp, mfld_simps] def proj : Z.total_space → B := bundle.total_space.proj @@ -510,7 +502,7 @@ end /-- Topological structure on the total space of a fiber bundle created from core, designed so that all the local trivialization are continuous. -/ -instance to_topological_space : topological_space (bundle.total_space Z.fiber) := +instance to_topological_space : topological_space Z.total_space := topological_space.generate_from $ ⋃ (i : ι) (s : set (B × F)) (s_open : is_open s), {(Z.local_triv_as_local_equiv i).source ∩ (Z.local_triv_as_local_equiv i) ⁻¹' s} @@ -567,7 +559,7 @@ def local_triv (i : ι) : trivialization F Z.proj := /-- Preferred local trivialization of a fiber bundle constructed from core, at a given point, as a bundle trivialization -/ -def local_triv_at (b : B) : trivialization F (π Z.fiber) := +def local_triv_at (b : B) : trivialization F (π F Z.fiber) := Z.local_triv (Z.index_at b) @[simp, mfld_simps] lemma local_triv_at_def (b : B) : @@ -647,11 +639,11 @@ by { rw [local_triv_at, ←base_set_at], exact Z.mem_base_set_at b, } /-- The inclusion of a fiber into the total space is a continuous map. -/ @[continuity] lemma continuous_total_space_mk (b : B) : - continuous (total_space_mk b : Z.fiber b → bundle.total_space Z.fiber) := + continuous (total_space.mk b : Z.fiber b → Z.total_space) := begin rw [continuous_iff_le_induced, fiber_bundle_core.to_topological_space], apply le_induced_generate_from, - simp only [total_space_mk, mem_Union, mem_singleton_iff, local_triv_as_local_equiv_source, + simp only [mem_Union, mem_singleton_iff, local_triv_as_local_equiv_source, local_triv_as_local_equiv_coe], rintros s ⟨i, t, ht, rfl⟩, rw [←((Z.local_triv i).source_inter_preimage_target_inter t), preimage_inter, ←preimage_comp, @@ -683,7 +675,6 @@ instance fiber_bundle : fiber_bundle F Z.fiber := (Z.local_triv_at b).open_source).mp (Z.local_triv_at b).continuous_to_fun _ ((Z.local_triv_at b).open_base_set.prod h), _⟩, rw [preimage_inter, ←preimage_comp, function.comp], - simp only [total_space_mk], refine ext_iff.mpr (λ a, ⟨λ ha, _, λ ha, ⟨Z.mem_base_set_at b, _⟩⟩), { simp only [mem_prod, mem_preimage, mem_inter_iff, local_triv_at_apply_mk] at ha, exact ha.2.2, }, @@ -713,21 +704,21 @@ topology in such a way that there is a fiber bundle structure for which the loca are also local homeomorphism and hence local trivializations. -/ @[nolint has_nonempty_instance] structure fiber_prebundle := -(pretrivialization_atlas : set (pretrivialization F (π E))) -(pretrivialization_at : B → pretrivialization F (π E)) +(pretrivialization_atlas : set (pretrivialization F (π F E))) +(pretrivialization_at : B → pretrivialization F (π F E)) (mem_base_pretrivialization_at : ∀ x : B, x ∈ (pretrivialization_at x).base_set) (pretrivialization_mem_atlas : ∀ x : B, pretrivialization_at x ∈ pretrivialization_atlas) (continuous_triv_change : ∀ e e' ∈ pretrivialization_atlas, continuous_on (e ∘ e'.to_local_equiv.symm) (e'.target ∩ (e'.to_local_equiv.symm ⁻¹' e.source))) -(total_space_mk_inducing : ∀ (b : B), inducing ((pretrivialization_at b) ∘ (total_space_mk b))) +(total_space_mk_inducing : ∀ (b : B), inducing ((pretrivialization_at b) ∘ (total_space.mk b))) namespace fiber_prebundle -variables {F E} (a : fiber_prebundle F E) {e : pretrivialization F (π E)} +variables {F E} (a : fiber_prebundle F E) {e : pretrivialization F (π F E)} /-- Topology on the total space that will make the prebundle into a bundle. -/ -def total_space_topology (a : fiber_prebundle F E) : topological_space (total_space E) := -⨆ (e : pretrivialization F (π E)) (he : e ∈ a.pretrivialization_atlas), +def total_space_topology (a : fiber_prebundle F E) : topological_space (total_space F E) := +⨆ (e : pretrivialization F (π F E)) (he : e ∈ a.pretrivialization_atlas), coinduced e.set_symm (subtype.topological_space) lemma continuous_symm_of_mem_pretrivialization_atlas (he : e ∈ a.pretrivialization_atlas) : @@ -739,7 +730,7 @@ begin exact le_supr₂ e he, end -lemma is_open_source (e : pretrivialization F (π E)) : is_open[a.total_space_topology] e.source := +lemma is_open_source (e : pretrivialization F (π F E)) : is_open[a.total_space_topology] e.source := begin letI := a.total_space_topology, refine is_open_supr_iff.mpr (λ e', _), @@ -751,7 +742,7 @@ begin pretrivialization.preimage_symm_proj_inter], end -lemma is_open_target_of_mem_pretrivialization_atlas_inter (e e' : pretrivialization F (π E)) +lemma is_open_target_of_mem_pretrivialization_atlas_inter (e e' : pretrivialization F (π F E)) (he' : e' ∈ a.pretrivialization_atlas) : is_open (e'.to_local_equiv.target ∩ e'.to_local_equiv.symm ⁻¹' e.source) := begin @@ -764,7 +755,7 @@ end /-- Promotion from a `pretrivialization` to a `trivialization`. -/ def trivialization_of_mem_pretrivialization_atlas (he : e ∈ a.pretrivialization_atlas) : - @trivialization B F _ _ _ a.total_space_topology (π E) := + @trivialization B F _ _ _ a.total_space_topology (π F E) := { open_source := a.is_open_source e, continuous_to_fun := begin letI := a.total_space_topology, @@ -785,14 +776,14 @@ def trivialization_of_mem_pretrivialization_atlas (he : e ∈ a.pretrivializatio .. e } lemma mem_trivialization_at_source (b : B) (x : E b) : - total_space_mk b x ∈ (a.pretrivialization_at b).source := + total_space.mk b x ∈ (a.pretrivialization_at b).source := begin simp only [(a.pretrivialization_at b).source_eq, mem_preimage, total_space.proj], exact a.mem_base_pretrivialization_at b, end @[simp] lemma total_space_mk_preimage_source (b : B) : - (total_space_mk b) ⁻¹' (a.pretrivialization_at b).source = univ := + total_space.mk b ⁻¹' (a.pretrivialization_at b).source = univ := begin apply eq_univ_of_univ_subset, rw [(a.pretrivialization_at b).source_eq, ←preimage_comp, function.comp], @@ -802,7 +793,7 @@ begin end @[continuity] lemma continuous_total_space_mk (b : B) : - @continuous _ _ _ a.total_space_topology (total_space_mk b) := + @continuous _ _ _ a.total_space_topology (total_space.mk b) := begin letI := a.total_space_topology, let e := a.trivialization_of_mem_pretrivialization_atlas (a.pretrivialization_mem_atlas b), @@ -812,8 +803,8 @@ begin end lemma inducing_total_space_mk_of_inducing_comp (b : B) - (h : inducing ((a.pretrivialization_at b) ∘ (total_space_mk b))) : - @inducing _ _ _ a.total_space_topology (total_space_mk b) := + (h : inducing ((a.pretrivialization_at b) ∘ (total_space.mk b))) : + @inducing _ _ _ a.total_space_topology (total_space.mk b) := begin letI := a.total_space_topology, rw ←restrict_comp_cod_restrict (a.mem_trivialization_at_source b) at h, @@ -841,7 +832,7 @@ def to_fiber_bundle : mem_base_set_trivialization_at := a.mem_base_pretrivialization_at, trivialization_mem_atlas := λ x, ⟨_, a.pretrivialization_mem_atlas x, rfl⟩ } -lemma continuous_proj : @continuous _ _ a.total_space_topology _ (π E) := +lemma continuous_proj : @continuous _ _ a.total_space_topology _ (π F E) := begin letI := a.total_space_topology, letI := a.to_fiber_bundle, @@ -849,17 +840,17 @@ begin end /-- For a fiber bundle `E` over `B` constructed using the `fiber_prebundle` mechanism, -continuity of a function `total_space E → X` on an open set `s` can be checked by precomposing at +continuity of a function `total_space F E → X` on an open set `s` can be checked by precomposing at each point with the pretrivialization used for the construction at that point. -/ -lemma continuous_on_of_comp_right {X : Type*} [topological_space X] {f : total_space E → X} +lemma continuous_on_of_comp_right {X : Type*} [topological_space X] {f : total_space F E → X} {s : set B} (hs : is_open s) (hf : ∀ b ∈ s, continuous_on (f ∘ (a.pretrivialization_at b).to_local_equiv.symm) ((s ∩ (a.pretrivialization_at b).base_set) ×ˢ (set.univ : set F))) : - @continuous_on _ _ a.total_space_topology _ f ((π E) ⁻¹' s) := + @continuous_on _ _ a.total_space_topology _ f ((π F E) ⁻¹' s) := begin letI := a.total_space_topology, intros z hz, - let e : trivialization F (π E) := + let e : trivialization F (π F E) := a.trivialization_of_mem_pretrivialization_atlas (a.pretrivialization_mem_atlas z.proj), refine (e.continuous_at_of_comp_right _ ((hf z.proj hz).continuous_at (is_open.mem_nhds _ _))).continuous_within_at, diff --git a/src/topology/fiber_bundle/constructions.lean b/src/topology/fiber_bundle/constructions.lean index 4607bf248af1c..a4d859c6166d3 100644 --- a/src/topology/fiber_bundle/constructions.lean +++ b/src/topology/fiber_bundle/constructions.lean @@ -37,23 +37,21 @@ namespace trivial variables (B : Type*) (F : Type*) -instance [I : topological_space F] : ∀ x : B, topological_space (trivial B F x) := λ x, I - instance [t₁ : topological_space B] [t₂ : topological_space F] : - topological_space (total_space (trivial B F)) := -induced total_space.proj t₁ ⊓ induced (trivial.proj_snd B F) t₂ + topological_space (total_space F (trivial B F)) := +induced total_space.proj t₁ ⊓ induced (total_space.trivial_snd B F) t₂ variables [topological_space B] [topological_space F] /-- Local trivialization for trivial bundle. -/ -def trivialization : trivialization F (π (bundle.trivial B F)) := -{ to_fun := λ x, (x.fst, x.snd), +def trivialization : trivialization F (π F (λ _ : B, F)) := +{ to_fun := λ x, (x.proj, x.snd), inv_fun := λ y, ⟨y.fst, y.snd⟩, source := univ, target := univ, - map_source' := λ x h, mem_univ (x.fst, x.snd), - map_target' := λ y h, mem_univ ⟨y.fst, y.snd⟩, - left_inv' := λ x h, sigma.eq rfl rfl, + map_source' := λ x h, mem_univ _, + map_target' := λ y h, mem_univ _, + left_inv' := λ x h, total_space.ext _ _ rfl heq.rfl, right_inv' := λ x h, prod.ext rfl rfl, open_source := is_open_univ, open_target := is_open_univ, @@ -83,11 +81,10 @@ instance fiber_bundle : fiber_bundle F (bundle.trivial B F) := total_space_mk_inducing := λ b, ⟨begin have : (λ (x : trivial B F b), x) = @id F, by { ext x, refl }, simp only [total_space.topological_space, induced_inf, induced_compose, function.comp, - total_space.proj, induced_const, top_inf_eq, trivial.proj_snd, id.def, - trivial.topological_space, this, induced_id], + induced_const, top_inf_eq, total_space.trivial_snd, id.def, this, induced_id], end⟩ } -lemma eq_trivialization (e : _root_.trivialization F (π (bundle.trivial B F))) +lemma eq_trivialization (e : _root_.trivialization F (π F (bundle.trivial B F))) [i : mem_trivialization_atlas e] : e = trivialization B F := i.out @@ -102,21 +99,22 @@ section prod variables {B : Type*} section defs -variables (E₁ : B → Type*) (E₂ : B → Type*) -variables [topological_space (total_space E₁)] [topological_space (total_space E₂)] +variables (F₁ : Type*) (E₁ : B → Type*) (F₂ : Type*) (E₂ : B → Type*) +variables [topological_space (total_space F₁ E₁)] [topological_space (total_space F₂ E₂)] /-- Equip the total space of the fiberwise product of two fiber bundles `E₁`, `E₂` with the induced topology from the diagonal embedding into `total_space E₁ × total_space E₂`. -/ -instance fiber_bundle.prod.topological_space : topological_space (total_space (E₁ ×ᵇ E₂)) := +instance fiber_bundle.prod.topological_space : + topological_space (total_space (F₁ × F₂) (E₁ ×ᵇ E₂)) := topological_space.induced - (λ p, ((⟨p.1, p.2.1⟩ : total_space E₁), (⟨p.1, p.2.2⟩ : total_space E₂))) - (by apply_instance : topological_space (total_space E₁ × total_space E₂)) + (λ p, ((⟨p.1, p.2.1⟩ : total_space F₁ E₁), (⟨p.1, p.2.2⟩ : total_space F₂ E₂))) + (by apply_instance : topological_space (total_space F₁ E₁ × total_space F₂ E₂)) /-- The diagonal map from the total space of the fiberwise product of two fiber bundles `E₁`, `E₂` into `total_space E₁ × total_space E₂` is `inducing`. -/ lemma fiber_bundle.prod.inducing_diag : inducing (λ p, (⟨p.1, p.2.1⟩, ⟨p.1, p.2.2⟩) : - total_space (E₁ ×ᵇ E₂) → total_space E₁ × total_space E₂) := + total_space (F₁ × F₂) (E₁ ×ᵇ E₂) → total_space F₁ E₁ × total_space F₂ E₂) := ⟨rfl⟩ end defs @@ -124,28 +122,28 @@ end defs open fiber_bundle variables [topological_space B] - (F₁ : Type*) [topological_space F₁] (E₁ : B → Type*) [topological_space (total_space E₁)] - (F₂ : Type*) [topological_space F₂] (E₂ : B → Type*) [topological_space (total_space E₂)] + (F₁ : Type*) [topological_space F₁] (E₁ : B → Type*) [topological_space (total_space F₁ E₁)] + (F₂ : Type*) [topological_space F₂] (E₂ : B → Type*) [topological_space (total_space F₂ E₂)] namespace trivialization -variables {F₁ E₁ F₂ E₂} (e₁ : trivialization F₁ (π E₁)) (e₂ : trivialization F₂ (π E₂)) +variables {F₁ E₁ F₂ E₂} (e₁ : trivialization F₁ (π F₁ E₁)) (e₂ : trivialization F₂ (π F₂ E₂)) /-- Given trivializations `e₁`, `e₂` for fiber bundles `E₁`, `E₂` over a base `B`, the forward function for the construction `trivialization.prod`, the induced trivialization for the fiberwise product of `E₁` and `E₂`. -/ -def prod.to_fun' : total_space (E₁ ×ᵇ E₂) → B × (F₁ × F₂) := +def prod.to_fun' : total_space (F₁ × F₂) (E₁ ×ᵇ E₂) → B × (F₁ × F₂) := λ p, ⟨p.1, (e₁ ⟨p.1, p.2.1⟩).2, (e₂ ⟨p.1, p.2.2⟩).2⟩ variables {e₁ e₂} -lemma prod.continuous_to_fun : continuous_on (prod.to_fun' e₁ e₂) - (@total_space.proj B (E₁ ×ᵇ E₂) ⁻¹' (e₁.base_set ∩ e₂.base_set)) := +lemma prod.continuous_to_fun : + continuous_on (prod.to_fun' e₁ e₂) (π (F₁ × F₂) (E₁ ×ᵇ E₂) ⁻¹' (e₁.base_set ∩ e₂.base_set)) := begin - let f₁ : total_space (E₁ ×ᵇ E₂) → total_space E₁ × total_space E₂ := - λ p, ((⟨p.1, p.2.1⟩ : total_space E₁), (⟨p.1, p.2.2⟩ : total_space E₂)), - let f₂ : total_space E₁ × total_space E₂ → (B × F₁) × (B × F₂) := λ p, ⟨e₁ p.1, e₂ p.2⟩, + let f₁ : total_space (F₁ × F₂) (E₁ ×ᵇ E₂) → total_space F₁ E₁ × total_space F₂ E₂ := + λ p, ((⟨p.1, p.2.1⟩ : total_space F₁ E₁), (⟨p.1, p.2.2⟩ : total_space F₂ E₂)), + let f₂ : total_space F₁ E₁ × total_space F₂ E₂ → (B × F₁) × (B × F₂) := λ p, ⟨e₁ p.1, e₂ p.2⟩, let f₃ : (B × F₁) × (B × F₂) → B × F₁ × F₂ := λ p, ⟨p.1.1, p.1.2, p.2.2⟩, - have hf₁ : continuous f₁ := (prod.inducing_diag E₁ E₂).continuous, + have hf₁ : continuous f₁ := (prod.inducing_diag F₁ E₁ F₂ E₂).continuous, have hf₂ : continuous_on f₂ (e₁.source ×ˢ e₂.source) := e₁.to_local_homeomorph.continuous_on.prod_map e₂.to_local_homeomorph.continuous_on, have hf₃ : continuous f₃ := @@ -165,18 +163,19 @@ variables (e₁ e₂) [Π x, has_zero (E₁ x)] [∀ x, has_zero (E₂ x)] /-- Given trivializations `e₁`, `e₂` for fiber bundles `E₁`, `E₂` over a base `B`, the inverse function for the construction `trivialization.prod`, the induced trivialization for the fiberwise product of `E₁` and `E₂`. -/ -noncomputable def prod.inv_fun' (p : B × (F₁ × F₂)) : total_space (E₁ ×ᵇ E₂) := +noncomputable def prod.inv_fun' (p : B × (F₁ × F₂)) : total_space (F₁ × F₂) (E₁ ×ᵇ E₂) := ⟨p.1, e₁.symm p.1 p.2.1, e₂.symm p.1 p.2.2⟩ variables {e₁ e₂} -lemma prod.left_inv {x : total_space (E₁ ×ᵇ E₂)} - (h : x ∈ @total_space.proj B (E₁ ×ᵇ E₂) ⁻¹' (e₁.base_set ∩ e₂.base_set)) : +lemma prod.left_inv {x : total_space (F₁ × F₂) (E₁ ×ᵇ E₂)} + (h : x ∈ π (F₁ × F₂) (E₁ ×ᵇ E₂) ⁻¹' (e₁.base_set ∩ e₂.base_set)) : prod.inv_fun' e₁ e₂ (prod.to_fun' e₁ e₂ x) = x := begin obtain ⟨x, v₁, v₂⟩ := x, obtain ⟨h₁ : x ∈ e₁.base_set, h₂ : x ∈ e₂.base_set⟩ := h, - simp only [prod.to_fun', prod.inv_fun', symm_apply_apply_mk, h₁, h₂] + simp only [prod.to_fun', prod.inv_fun', symm_apply_apply_mk, h₁, h₂, + eq_self_iff_true, heq_iff_eq, and_self] end lemma prod.right_inv {x : B × F₁ × F₂} @@ -191,22 +190,22 @@ end lemma prod.continuous_inv_fun : continuous_on (prod.inv_fun' e₁ e₂) ((e₁.base_set ∩ e₂.base_set) ×ˢ univ) := begin - rw (prod.inducing_diag E₁ E₂).continuous_on_iff, + rw (prod.inducing_diag F₁ E₁ F₂ E₂).continuous_on_iff, have H₁ : continuous (λ p : B × F₁ × F₂, ((p.1, p.2.1), (p.1, p.2.2))) := (continuous_id.prod_map continuous_fst).prod_mk (continuous_id.prod_map continuous_snd), refine (e₁.continuous_on_symm.prod_map e₂.continuous_on_symm).comp H₁.continuous_on _, exact λ x h, ⟨⟨h.1.1, mem_univ _⟩, ⟨h.1.2, mem_univ _⟩⟩ end -variables (e₁ e₂ e₁ e₂) +variables (e₁ e₂) /-- Given trivializations `e₁`, `e₂` for bundle types `E₁`, `E₂` over a base `B`, the induced trivialization for the fiberwise product of `E₁` and `E₂`, whose base set is `e₁.base_set ∩ e₂.base_set`. -/ -noncomputable def prod : trivialization (F₁ × F₂) (π (E₁ ×ᵇ E₂)) := +noncomputable def prod : trivialization (F₁ × F₂) (π (F₁ × F₂) (E₁ ×ᵇ E₂)) := { to_fun := prod.to_fun' e₁ e₂, inv_fun := prod.inv_fun' e₁ e₂, - source := (@total_space.proj B (E₁ ×ᵇ E₂)) ⁻¹' (e₁.base_set ∩ e₂.base_set), + source := (π (F₁ × F₂) (E₁ ×ᵇ E₂)) ⁻¹' (e₁.base_set ∩ e₂.base_set), target := (e₁.base_set ∩ e₂.base_set) ×ˢ set.univ, map_source' := λ x h, ⟨h, set.mem_univ _⟩, map_target' := λ x h, h.1, @@ -214,7 +213,7 @@ noncomputable def prod : trivialization (F₁ × F₂) (π (E₁ ×ᵇ E₂)) := right_inv' := λ x, prod.right_inv, open_source := begin convert (e₁.open_source.prod e₂.open_source).preimage - (fiber_bundle.prod.inducing_diag E₁ E₂).continuous, + (fiber_bundle.prod.inducing_diag F₁ E₁ F₂ E₂).continuous, ext x, simp only [trivialization.source_eq] with mfld_simps, end, @@ -230,8 +229,8 @@ noncomputable def prod : trivialization (F₁ × F₂) (π (E₁ ×ᵇ E₂)) := @[simp] lemma base_set_prod : (prod e₁ e₂).base_set = e₁.base_set ∩ e₂.base_set := rfl -lemma prod_symm_apply (x : B) (w₁ : F₁) (w₂ : F₂) : (prod e₁ e₂).to_local_equiv.symm (x, w₁, w₂) - = ⟨x, e₁.symm x w₁, e₂.symm x w₂⟩ := +lemma prod_symm_apply (x : B) (w₁ : F₁) (w₂ : F₂) : + (prod e₁ e₂).to_local_equiv.symm (x, w₁, w₂) = ⟨x, e₁.symm x w₁, e₂.symm x w₂⟩ := rfl end trivialization @@ -246,11 +245,11 @@ variables [Π x, has_zero (E₁ x)] [∀ x, has_zero (E₂ x)] noncomputable instance fiber_bundle.prod : fiber_bundle (F₁ × F₂) (E₁ ×ᵇ E₂) := { total_space_mk_inducing := λ b, begin - rw (prod.inducing_diag E₁ E₂).inducing_iff, + rw (prod.inducing_diag F₁ E₁ F₂ E₂).inducing_iff, exact (total_space_mk_inducing F₁ E₁ b).prod_mk (total_space_mk_inducing F₂ E₂ b), end, trivialization_atlas := - {e | ∃ (e₁ : trivialization F₁ (π E₁)) (e₂ : trivialization F₂ (π E₂)) + {e | ∃ (e₁ : trivialization F₁ (π F₁ E₁)) (e₂ : trivialization F₂ (π F₂ E₂)) [mem_trivialization_atlas e₁] [mem_trivialization_atlas e₂], by exactI e = trivialization.prod e₁ e₂}, trivialization_at := λ b, (trivialization_at F₁ E₁ b).prod (trivialization_at F₂ E₂ b), @@ -259,9 +258,9 @@ noncomputable instance fiber_bundle.prod : fiber_bundle (F₁ × F₂) (E₁ × trivialization_mem_atlas := λ b, ⟨trivialization_at F₁ E₁ b, trivialization_at F₂ E₂ b, by apply_instance, by apply_instance, rfl⟩ } -instance {e₁ : trivialization F₁ (π E₁)} {e₂ : trivialization F₂ (π E₂)} +instance {e₁ : trivialization F₁ (π F₁ E₁)} {e₂ : trivialization F₂ (π F₂ E₂)} [mem_trivialization_atlas e₁] [mem_trivialization_atlas e₂] : - mem_trivialization_atlas (e₁.prod e₂ : trivialization (F₁ × F₂) (π (E₁ ×ᵇ E₂))) := + mem_trivialization_atlas (e₁.prod e₂ : trivialization (F₁ × F₂) (π (F₁ × F₂) (E₁ ×ᵇ E₂))) := { out := ⟨e₁, e₂, by apply_instance, by apply_instance, rfl⟩ } end prod @@ -274,32 +273,32 @@ variables {B : Type*} (F : Type*) (E : B → Type*) {B' : Type*} (f : B' → B) instance [∀ (x : B), topological_space (E x)] : ∀ (x : B'), topological_space ((f *ᵖ E) x) := by delta_instance bundle.pullback -variables [topological_space B'] [topological_space (total_space E)] +variables [topological_space B'] [topological_space (total_space F E)] /-- Definition of `pullback.total_space.topological_space`, which we make irreducible. -/ -@[irreducible] def pullback_topology : topological_space (total_space (f *ᵖ E)) := +@[irreducible] def pullback_topology : topological_space (total_space F (f *ᵖ E)) := induced total_space.proj ‹topological_space B'› ⊓ -induced (pullback.lift f) ‹topological_space (total_space E)› +induced (pullback.lift f) ‹topological_space (total_space F E)› /-- The topology on the total space of a pullback bundle is the coarsest topology for which both the projections to the base and the map to the original bundle are continuous. -/ -instance pullback.total_space.topological_space : topological_space (total_space (f *ᵖ E)) := -pullback_topology E f +instance pullback.total_space.topological_space : topological_space (total_space F (f *ᵖ E)) := +pullback_topology F E f -lemma pullback.continuous_proj (f : B' → B) : continuous (@total_space.proj _ (f *ᵖ E)) := +lemma pullback.continuous_proj (f : B' → B) : continuous (π F (f *ᵖ E)) := begin rw [continuous_iff_le_induced, pullback.total_space.topological_space, pullback_topology], exact inf_le_left, end -lemma pullback.continuous_lift (f : B' → B) : continuous (@pullback.lift B E B' f) := +lemma pullback.continuous_lift (f : B' → B) : continuous (@pullback.lift B F E B' f) := begin rw [continuous_iff_le_induced, pullback.total_space.topological_space, pullback_topology], exact inf_le_right, end lemma inducing_pullback_total_space_embedding (f : B' → B) : - inducing (@pullback_total_space_embedding B E B' f) := + inducing (@pullback_total_space_embedding B F E B' f) := begin constructor, simp_rw [prod.topological_space, induced_inf, induced_compose, @@ -312,44 +311,42 @@ variables (F) [topological_space F] [topological_space B] lemma pullback.continuous_total_space_mk [∀ x, topological_space (E x)] [fiber_bundle F E] {f : B' → B} {x : B'} : - continuous (@total_space_mk _ (f *ᵖ E) x) := + continuous (@total_space.mk _ F (f *ᵖ E) x) := begin simp only [continuous_iff_le_induced, pullback.total_space.topological_space, induced_compose, - induced_inf, function.comp, total_space_mk, total_space.proj, induced_const, top_inf_eq, - pullback_topology], + induced_inf, function.comp, induced_const, top_inf_eq, pullback_topology], exact le_of_eq (fiber_bundle.total_space_mk_inducing F E (f x)).induced, end variables {E F} [∀ b, has_zero (E b)] {K : Type*} [continuous_map_class K B' B] /-- A fiber bundle trivialization can be pulled back to a trivialization on the pullback bundle. -/ -noncomputable def trivialization.pullback (e : trivialization F (π E)) (f : K) : - trivialization F (π ((f : B' → B) *ᵖ E)) := +noncomputable def trivialization.pullback (e : trivialization F (π F E)) (f : K) : + trivialization F (π F ((f : B' → B) *ᵖ E)) := { to_fun := λ z, (z.proj, (e (pullback.lift f z)).2), - inv_fun := λ y, @total_space_mk _ (f *ᵖ E) y.1 (e.symm (f y.1) y.2), + inv_fun := λ y, @total_space.mk _ _ (f *ᵖ E) y.1 (e.symm (f y.1) y.2), source := pullback.lift f ⁻¹' e.source, base_set := f ⁻¹' e.base_set, target := (f ⁻¹' e.base_set) ×ˢ univ, - map_source' := λ x h, by { simp_rw [e.source_eq, mem_preimage, pullback.proj_lift] at h, + map_source' := λ x h, by { simp_rw [e.source_eq, mem_preimage, pullback.lift_proj] at h, simp_rw [prod_mk_mem_set_prod_eq, mem_univ, and_true, mem_preimage, h] }, map_target' := λ y h, by { rw [mem_prod, mem_preimage] at h, - simp_rw [e.source_eq, mem_preimage, pullback.proj_lift, h.1] }, - left_inv' := λ x h, by { simp_rw [mem_preimage, e.mem_source, pullback.proj_lift] at h, + simp_rw [e.source_eq, mem_preimage, pullback.lift_proj, h.1] }, + left_inv' := λ x h, by { simp_rw [mem_preimage, e.mem_source, pullback.lift_proj] at h, simp_rw [pullback.lift, e.symm_apply_apply_mk h, total_space.eta] }, right_inv' := λ x h, by { simp_rw [mem_prod, mem_preimage, mem_univ, and_true] at h, - simp_rw [total_space.proj_mk, pullback.lift_mk, e.apply_mk_symm h, prod.mk.eta] }, + simp_rw [pullback.lift_mk, e.apply_mk_symm h, prod.mk.eta] }, open_source := by { simp_rw [e.source_eq, ← preimage_comp], exact ((map_continuous f).comp $ - pullback.continuous_proj E f).is_open_preimage _ e.open_base_set }, + pullback.continuous_proj F E f).is_open_preimage _ e.open_base_set }, open_target := ((map_continuous f).is_open_preimage _ e.open_base_set).prod is_open_univ, open_base_set := (map_continuous f).is_open_preimage _ e.open_base_set, - continuous_to_fun := (pullback.continuous_proj E f).continuous_on.prod + continuous_to_fun := (pullback.continuous_proj F E f).continuous_on.prod (continuous_snd.comp_continuous_on $ - e.continuous_on.comp (pullback.continuous_lift E f).continuous_on subset.rfl), + e.continuous_on.comp (pullback.continuous_lift F E f).continuous_on subset.rfl), continuous_inv_fun := begin dsimp only, - simp_rw [(inducing_pullback_total_space_embedding E f).continuous_on_iff, function.comp, - pullback_total_space_embedding, total_space.proj_mk], - dsimp only [total_space.proj_mk], + simp_rw [(inducing_pullback_total_space_embedding F E f).continuous_on_iff, function.comp, + pullback_total_space_embedding], refine continuous_on_fst.prod (e.continuous_on_symm.comp ((map_continuous f).prod_map continuous_id).continuous_on subset.rfl) end, @@ -360,10 +357,10 @@ noncomputable def trivialization.pullback (e : trivialization F (π E)) (f : K) noncomputable instance fiber_bundle.pullback [∀ x, topological_space (E x)] [fiber_bundle F E] (f : K) : fiber_bundle F ((f : B' → B) *ᵖ E) := { total_space_mk_inducing := λ x, inducing_of_inducing_compose - (pullback.continuous_total_space_mk F E) (pullback.continuous_lift E f) + (pullback.continuous_total_space_mk F E) (pullback.continuous_lift F E f) (total_space_mk_inducing F E (f x)), trivialization_atlas := - {ef | ∃ (e : trivialization F (π E)) [mem_trivialization_atlas e], ef = e.pullback f}, + {ef | ∃ (e : trivialization F (π F E)) [mem_trivialization_atlas e], ef = e.pullback f}, trivialization_at := λ x, (trivialization_at F E (f x)).pullback f, mem_base_set_trivialization_at := λ x, mem_base_set_trivialization_at F E (f x), trivialization_mem_atlas := λ x, ⟨trivialization_at F E (f x), by apply_instance, rfl⟩ } diff --git a/src/topology/fiber_bundle/trivialization.lean b/src/topology/fiber_bundle/trivialization.lean index 84a15ebb82f14..fcde8b14a33e4 100644 --- a/src/topology/fiber_bundle/trivialization.lean +++ b/src/topology/fiber_bundle/trivialization.lean @@ -159,9 +159,9 @@ lemma symm_trans_target_eq (e e' : pretrivialization F proj) : (e.to_local_equiv.symm.trans e'.to_local_equiv).target = (e.base_set ∩ e'.base_set) ×ˢ univ := by rw [← local_equiv.symm_source, symm_trans_symm, symm_trans_source_eq, inter_comm] -variables {B F} (e' : pretrivialization F (π E)) {x' : total_space E} {b : B} {y : E b} +variables {B F} (e' : pretrivialization F (π F E)) {x' : total_space F E} {b : B} {y : E b} -lemma coe_mem_source : ↑y ∈ e'.source ↔ b ∈ e'.base_set := e'.mem_source +@[simp] theorem coe_mem_source : ↑y ∈ e'.source ↔ b ∈ e'.base_set := e'.mem_source @[simp, mfld_simps] lemma coe_coe_fst (hb : b ∈ e'.base_set) : (e' y).1 = b := e'.coe_fst (e'.mem_source.2 hb) @@ -169,7 +169,7 @@ e'.coe_fst (e'.mem_source.2 hb) lemma mk_mem_target {x : B} {y : F} : (x, y) ∈ e'.target ↔ x ∈ e'.base_set := e'.mem_target -lemma symm_coe_proj {x : B} {y : F} (e' : pretrivialization F (π E)) (h : x ∈ e'.base_set) : +lemma symm_coe_proj {x : B} {y : F} (e' : pretrivialization F (π F E)) (h : x ∈ e'.base_set) : (e'.to_local_equiv.symm (x, y)).1 = x := e'.proj_symm_apply' h @@ -177,46 +177,46 @@ section has_zero variables [∀ x, has_zero (E x)] /-- A fiberwise inverse to `e`. This is the function `F → E b` that induces a local inverse -`B × F → total_space E` of `e` on `e.base_set`. It is defined to be `0` outside `e.base_set`. -/ -protected noncomputable def symm (e : pretrivialization F (π E)) (b : B) (y : F) : E b := +`B × F → total_space F E` of `e` on `e.base_set`. It is defined to be `0` outside `e.base_set`. -/ +protected noncomputable def symm (e : pretrivialization F (π F E)) (b : B) (y : F) : E b := if hb : b ∈ e.base_set then cast (congr_arg E (e.proj_symm_apply' hb)) (e.to_local_equiv.symm (b, y)).2 else 0 -lemma symm_apply (e : pretrivialization F (π E)) {b : B} (hb : b ∈ e.base_set) (y : F) : +lemma symm_apply (e : pretrivialization F (π F E)) {b : B} (hb : b ∈ e.base_set) (y : F) : e.symm b y = cast (congr_arg E (e.symm_coe_proj hb)) (e.to_local_equiv.symm (b, y)).2 := dif_pos hb -lemma symm_apply_of_not_mem (e : pretrivialization F (π E)) {b : B} (hb : b ∉ e.base_set) (y : F) : - e.symm b y = 0 := +lemma symm_apply_of_not_mem (e : pretrivialization F (π F E)) {b : B} (hb : b ∉ e.base_set) + (y : F) : e.symm b y = 0 := dif_neg hb -lemma coe_symm_of_not_mem (e : pretrivialization F (π E)) {b : B} (hb : b ∉ e.base_set) : +lemma coe_symm_of_not_mem (e : pretrivialization F (π F E)) {b : B} (hb : b ∉ e.base_set) : (e.symm b : F → E b) = 0 := funext $ λ y, dif_neg hb -lemma mk_symm (e : pretrivialization F (π E)) {b : B} (hb : b ∈ e.base_set) (y : F) : - total_space_mk b (e.symm b y) = e.to_local_equiv.symm (b, y) := +lemma mk_symm (e : pretrivialization F (π F E)) {b : B} (hb : b ∈ e.base_set) (y : F) : + total_space.mk b (e.symm b y) = e.to_local_equiv.symm (b, y) := by rw [e.symm_apply hb, total_space.mk_cast, total_space.eta] -lemma symm_proj_apply (e : pretrivialization F (π E)) (z : total_space E) +lemma symm_proj_apply (e : pretrivialization F (π F E)) (z : total_space F E) (hz : z.proj ∈ e.base_set) : e.symm z.proj (e z).2 = z.2 := by rw [e.symm_apply hz, cast_eq_iff_heq, e.mk_proj_snd' hz, e.symm_apply_apply (e.mem_source.mpr hz)] -lemma symm_apply_apply_mk (e : pretrivialization F (π E)) {b : B} (hb : b ∈ e.base_set) (y : E b) : - e.symm b (e (total_space_mk b y)).2 = y := -e.symm_proj_apply (total_space_mk b y) hb +lemma symm_apply_apply_mk (e : pretrivialization F (π F E)) {b : B} (hb : b ∈ e.base_set) + (y : E b) : e.symm b (e ⟨b, y⟩).2 = y := +e.symm_proj_apply ⟨b, y⟩ hb -lemma apply_mk_symm (e : pretrivialization F (π E)) {b : B} (hb : b ∈ e.base_set) (y : F) : - e (total_space_mk b (e.symm b y)) = (b, y) := +lemma apply_mk_symm (e : pretrivialization F (π F E)) {b : B} (hb : b ∈ e.base_set) (y : F) : + e ⟨b, e.symm b y⟩ = (b, y) := by rw [e.mk_symm hb, e.apply_symm_apply (e.mk_mem_target.mpr hb)] end has_zero end pretrivialization -variables [topological_space Z] [topological_space (total_space E)] +variables [topological_space Z] [topological_space (total_space F E)] /-- A structure extending local homeomorphisms, defining a local trivialization of a projection @@ -403,7 +403,7 @@ begin exact hf_proj.preimage_mem_nhds (e.open_base_set.mem_nhds he), end -variables {E} (e' : trivialization F (π E)) {x' : total_space E} {b : B} {y : E b} +variables {E} (e' : trivialization F (π F E)) {x' : total_space F E} {b : B} {y : E b} protected lemma continuous_on : continuous_on e' e'.source := e'.continuous_to_fun @@ -418,51 +418,52 @@ e'.coe_fst (e'.mem_source.2 hb) lemma mk_mem_target {y : F} : (b, y) ∈ e'.target ↔ b ∈ e'.base_set := e'.to_pretrivialization.mem_target -lemma symm_apply_apply {x : total_space E} (hx : x ∈ e'.source) : +lemma symm_apply_apply {x : total_space F E} (hx : x ∈ e'.source) : e'.to_local_homeomorph.symm (e' x) = x := e'.to_local_equiv.left_inv hx @[simp, mfld_simps] lemma symm_coe_proj {x : B} {y : F} - (e : trivialization F (π E)) (h : x ∈ e.base_set) : + (e : trivialization F (π F E)) (h : x ∈ e.base_set) : (e.to_local_homeomorph.symm (x, y)).1 = x := e.proj_symm_apply' h section has_zero variables [∀ x, has_zero (E x)] /-- A fiberwise inverse to `e'`. The function `F → E x` that induces a local inverse -`B × F → total_space E` of `e'` on `e'.base_set`. It is defined to be `0` outside `e'.base_set`. -/ -protected noncomputable def symm (e : trivialization F (π E)) (b : B) (y : F) : E b := +`B × F → total_space F E` of `e'` on `e'.base_set`. It is defined to be `0` outside +`e'.base_set`. -/ +protected noncomputable def symm (e : trivialization F (π F E)) (b : B) (y : F) : E b := e.to_pretrivialization.symm b y -lemma symm_apply (e : trivialization F (π E)) {b : B} (hb : b ∈ e.base_set) (y : F) : +lemma symm_apply (e : trivialization F (π F E)) {b : B} (hb : b ∈ e.base_set) (y : F) : e.symm b y = cast (congr_arg E (e.symm_coe_proj hb)) (e.to_local_homeomorph.symm (b, y)).2 := dif_pos hb -lemma symm_apply_of_not_mem (e : trivialization F (π E)) {b : B} (hb : b ∉ e.base_set) (y : F) : +lemma symm_apply_of_not_mem (e : trivialization F (π F E)) {b : B} (hb : b ∉ e.base_set) (y : F) : e.symm b y = 0 := dif_neg hb -lemma mk_symm (e : trivialization F (π E)) {b : B} (hb : b ∈ e.base_set) (y : F) : - total_space_mk b (e.symm b y) = e.to_local_homeomorph.symm (b, y) := +lemma mk_symm (e : trivialization F (π F E)) {b : B} (hb : b ∈ e.base_set) (y : F) : + total_space.mk b (e.symm b y) = e.to_local_homeomorph.symm (b, y) := e.to_pretrivialization.mk_symm hb y -lemma symm_proj_apply (e : trivialization F (π E)) (z : total_space E) +lemma symm_proj_apply (e : trivialization F (π F E)) (z : total_space F E) (hz : z.proj ∈ e.base_set) : e.symm z.proj (e z).2 = z.2 := e.to_pretrivialization.symm_proj_apply z hz -lemma symm_apply_apply_mk (e : trivialization F (π E)) {b : B} (hb : b ∈ e.base_set) (y : E b) : - e.symm b (e (total_space_mk b y)).2 = y := -e.symm_proj_apply (total_space_mk b y) hb +lemma symm_apply_apply_mk (e : trivialization F (π F E)) {b : B} (hb : b ∈ e.base_set) (y : E b) : + e.symm b (e ⟨b, y⟩).2 = y := +e.symm_proj_apply ⟨b, y⟩ hb -lemma apply_mk_symm (e : trivialization F (π E)) {b : B} (hb : b ∈ e.base_set) (y : F) : - e (total_space_mk b (e.symm b y)) = (b, y) := +lemma apply_mk_symm (e : trivialization F (π F E)) {b : B} (hb : b ∈ e.base_set) (y : F) : + e ⟨b, e.symm b y⟩ = (b, y) := e.to_pretrivialization.apply_mk_symm hb y -lemma continuous_on_symm (e : trivialization F (π E)) : - continuous_on (λ z : B × F, total_space_mk z.1 (e.symm z.1 z.2)) (e.base_set ×ˢ univ) := +lemma continuous_on_symm (e : trivialization F (π F E)) : + continuous_on (λ z : B × F, total_space.mk' F z.1 (e.symm z.1 z.2)) (e.base_set ×ˢ univ) := begin have : ∀ (z : B × F) (hz : z ∈ e.base_set ×ˢ (univ : set F)), - total_space_mk z.1 (e.symm z.1 z.2) = e.to_local_homeomorph.symm z, + total_space.mk z.1 (e.symm z.1 z.2) = e.to_local_homeomorph.symm z, { rintro x ⟨hx : x.1 ∈ e.base_set, _⟩, simp_rw [e.mk_symm hx, prod.mk.eta] }, refine continuous_on.congr _ this, rw [← e.target_eq], diff --git a/src/topology/vector_bundle/basic.lean b/src/topology/vector_bundle/basic.lean index 6d1d563cfcd6c..5a703818a8660 100644 --- a/src/topology/vector_bundle/basic.lean +++ b/src/topology/vector_bundle/basic.lean @@ -19,7 +19,7 @@ Let `B` be the base space, let `F` be a normed space over a normed field `R`, an `E : B → Type*` be a `fiber_bundle` with fiber `F`, in which, for each `x`, the fiber `E x` is a topological vector space over `R`. -To have a vector bundle structure on `bundle.total_space E`, one should additionally have the +To have a vector bundle structure on `bundle.total_space F E`, one should additionally have the following properties: * The bundle trivializations in the trivialization atlas should be continuous linear equivs in the @@ -70,23 +70,23 @@ variables {B F E} [semiring R] /-- A mixin class for `pretrivialization`, stating that a pretrivialization is fiberwise linear with respect to given module structures on its fibers and the model fiber. -/ protected class pretrivialization.is_linear [add_comm_monoid F] [module R F] - [∀ x, add_comm_monoid (E x)] [∀ x, module R (E x)] (e : pretrivialization F (π E)) : + [∀ x, add_comm_monoid (E x)] [∀ x, module R (E x)] (e : pretrivialization F (π F E)) : Prop := -(linear : ∀ b ∈ e.base_set, is_linear_map R (λ x : E b, (e (total_space_mk b x)).2)) +(linear : ∀ b ∈ e.base_set, is_linear_map R (λ x : E b, (e ⟨b, x⟩).2)) namespace pretrivialization -variables {F E} (e : pretrivialization F (π E)) {x : total_space E} {b : B} {y : E b} +variables {F E} (e : pretrivialization F (π F E)) {x : total_space F E} {b : B} {y : E b} lemma linear [add_comm_monoid F] [module R F] [∀ x, add_comm_monoid (E x)] [∀ x, module R (E x)] [e.is_linear R] {b : B} (hb : b ∈ e.base_set) : - is_linear_map R (λ x : E b, (e (total_space_mk b x)).2) := + is_linear_map R (λ x : E b, (e ⟨b, x⟩).2) := pretrivialization.is_linear.linear b hb variables [add_comm_monoid F] [module R F] [∀ x, add_comm_monoid (E x)] [∀ x, module R (E x)] /-- A fiberwise linear inverse to `e`. -/ -@[simps] protected def symmₗ (e : pretrivialization F (π E)) [e.is_linear R] (b : B) : +@[simps] protected def symmₗ (e : pretrivialization F (π F E)) [e.is_linear R] (b : B) : F →ₗ[R] E b := begin refine is_linear_map.mk' (e.symm b) _, @@ -98,10 +98,10 @@ end /-- A pretrivialization for a vector bundle defines linear equivalences between the fibers and the model space. -/ -@[simps {fully_applied := ff}] def linear_equiv_at (e : pretrivialization F (π E)) [e.is_linear R] +@[simps {fully_applied := ff}] def linear_equiv_at (e : pretrivialization F (π F E)) [e.is_linear R] (b : B) (hb : b ∈ e.base_set) : E b ≃ₗ[R] F := -{ to_fun := λ y, (e (total_space_mk b y)).2, +{ to_fun := λ y, (e ⟨b, y⟩).2, inv_fun := e.symm b, left_inv := e.symm_apply_apply_mk hb, right_inv := λ v, by simp_rw [e.apply_mk_symm hb v], @@ -109,66 +109,67 @@ fibers and the model space. -/ map_smul' := λ c v, (e.linear R hb).map_smul c v } /-- A fiberwise linear map equal to `e` on `e.base_set`. -/ -protected def linear_map_at (e : pretrivialization F (π E)) [e.is_linear R] (b : B) : E b →ₗ[R] F := +protected def linear_map_at (e : pretrivialization F (π F E)) [e.is_linear R] (b : B) : + E b →ₗ[R] F := if hb : b ∈ e.base_set then e.linear_equiv_at R b hb else 0 variables {R} -lemma coe_linear_map_at (e : pretrivialization F (π E)) [e.is_linear R] (b : B) : - ⇑(e.linear_map_at R b) = λ y, if b ∈ e.base_set then (e (total_space_mk b y)).2 else 0 := +lemma coe_linear_map_at (e : pretrivialization F (π F E)) [e.is_linear R] (b : B) : + ⇑(e.linear_map_at R b) = λ y, if b ∈ e.base_set then (e ⟨b, y⟩).2 else 0 := by { rw [pretrivialization.linear_map_at], split_ifs; refl } -lemma coe_linear_map_at_of_mem (e : pretrivialization F (π E)) [e.is_linear R] {b : B} +lemma coe_linear_map_at_of_mem (e : pretrivialization F (π F E)) [e.is_linear R] {b : B} (hb : b ∈ e.base_set) : - ⇑(e.linear_map_at R b) = λ y, (e (total_space_mk b y)).2 := + ⇑(e.linear_map_at R b) = λ y, (e ⟨b, y⟩).2 := by simp_rw [coe_linear_map_at, if_pos hb] -lemma linear_map_at_apply (e : pretrivialization F (π E)) [e.is_linear R] {b : B} (y : E b) : - e.linear_map_at R b y = if b ∈ e.base_set then (e (total_space_mk b y)).2 else 0 := +lemma linear_map_at_apply (e : pretrivialization F (π F E)) [e.is_linear R] {b : B} (y : E b) : + e.linear_map_at R b y = if b ∈ e.base_set then (e ⟨b, y⟩).2 else 0 := by rw [coe_linear_map_at] -lemma linear_map_at_def_of_mem (e : pretrivialization F (π E)) [e.is_linear R] {b : B} +lemma linear_map_at_def_of_mem (e : pretrivialization F (π F E)) [e.is_linear R] {b : B} (hb : b ∈ e.base_set) : e.linear_map_at R b = e.linear_equiv_at R b hb := dif_pos hb -lemma linear_map_at_def_of_not_mem (e : pretrivialization F (π E)) [e.is_linear R] {b : B} +lemma linear_map_at_def_of_not_mem (e : pretrivialization F (π F E)) [e.is_linear R] {b : B} (hb : b ∉ e.base_set) : e.linear_map_at R b = 0 := dif_neg hb -lemma linear_map_at_eq_zero (e : pretrivialization F (π E)) [e.is_linear R] {b : B} +lemma linear_map_at_eq_zero (e : pretrivialization F (π F E)) [e.is_linear R] {b : B} (hb : b ∉ e.base_set) : e.linear_map_at R b = 0 := dif_neg hb -lemma symmₗ_linear_map_at (e : pretrivialization F (π E)) [e.is_linear R] {b : B} +lemma symmₗ_linear_map_at (e : pretrivialization F (π F E)) [e.is_linear R] {b : B} (hb : b ∈ e.base_set) (y : E b) : e.symmₗ R b (e.linear_map_at R b y) = y := by { rw [e.linear_map_at_def_of_mem hb], exact (e.linear_equiv_at R b hb).left_inv y } -lemma linear_map_at_symmₗ (e : pretrivialization F (π E)) [e.is_linear R] {b : B} +lemma linear_map_at_symmₗ (e : pretrivialization F (π F E)) [e.is_linear R] {b : B} (hb : b ∈ e.base_set) (y : F) : e.linear_map_at R b (e.symmₗ R b y) = y := by { rw [e.linear_map_at_def_of_mem hb], exact (e.linear_equiv_at R b hb).right_inv y } end pretrivialization -variables (R) [topological_space (total_space E)] +variables (R) [topological_space (total_space F E)] /-- A mixin class for `trivialization`, stating that a trivialization is fiberwise linear with respect to given module structures on its fibers and the model fiber. -/ protected class trivialization.is_linear [add_comm_monoid F] [module R F] - [∀ x, add_comm_monoid (E x)] [∀ x, module R (E x)] (e : trivialization F (π E)) : Prop := -(linear : ∀ b ∈ e.base_set, is_linear_map R (λ x : E b, (e (total_space_mk b x)).2)) + [∀ x, add_comm_monoid (E x)] [∀ x, module R (E x)] (e : trivialization F (π F E)) : Prop := +(linear : ∀ b ∈ e.base_set, is_linear_map R (λ x : E b, (e ⟨b, x⟩).2)) namespace trivialization -variables (e : trivialization F (π E)) {x : total_space E} {b : B} {y : E b} +variables (e : trivialization F (π F E)) {x : total_space F E} {b : B} {y : E b} protected lemma linear [add_comm_monoid F] [module R F] [∀ x, add_comm_monoid (E x)] [∀ x, module R (E x)] [e.is_linear R] {b : B} (hb : b ∈ e.base_set) : - is_linear_map R (λ y : E b, (e (total_space_mk b y)).2) := + is_linear_map R (λ y : E b, (e ⟨b, y⟩).2) := trivialization.is_linear.linear b hb instance to_pretrivialization.is_linear [add_comm_monoid F] [module R F] @@ -180,71 +181,72 @@ variables [add_comm_monoid F] [module R F] [∀ x, add_comm_monoid (E x)] [∀ x /-- A trivialization for a vector bundle defines linear equivalences between the fibers and the model space. -/ -def linear_equiv_at (e : trivialization F (π E)) [e.is_linear R] (b : B) (hb : b ∈ e.base_set) : +def linear_equiv_at (e : trivialization F (π F E)) [e.is_linear R] (b : B) (hb : b ∈ e.base_set) : E b ≃ₗ[R] F := e.to_pretrivialization.linear_equiv_at R b hb variables {R} @[simp] -lemma linear_equiv_at_apply (e : trivialization F (π E)) [e.is_linear R] (b : B) +lemma linear_equiv_at_apply (e : trivialization F (π F E)) [e.is_linear R] (b : B) (hb : b ∈ e.base_set) (v : E b) : - e.linear_equiv_at R b hb v = (e (total_space_mk b v)).2 := rfl + e.linear_equiv_at R b hb v = (e ⟨b, v⟩).2 := rfl @[simp] -lemma linear_equiv_at_symm_apply (e : trivialization F (π E)) [e.is_linear R] (b : B) +lemma linear_equiv_at_symm_apply (e : trivialization F (π F E)) [e.is_linear R] (b : B) (hb : b ∈ e.base_set) (v : F) : (e.linear_equiv_at R b hb).symm v = e.symm b v := rfl variables (R) /-- A fiberwise linear inverse to `e`. -/ -protected def symmₗ (e : trivialization F (π E)) [e.is_linear R] (b : B) : F →ₗ[R] E b := +protected def symmₗ (e : trivialization F (π F E)) [e.is_linear R] (b : B) : F →ₗ[R] E b := e.to_pretrivialization.symmₗ R b variables {R} -lemma coe_symmₗ (e : trivialization F (π E)) [e.is_linear R] (b : B) : ⇑(e.symmₗ R b) = e.symm b := +lemma coe_symmₗ (e : trivialization F (π F E)) [e.is_linear R] (b : B) : + ⇑(e.symmₗ R b) = e.symm b := rfl variables (R) /-- A fiberwise linear map equal to `e` on `e.base_set`. -/ -protected def linear_map_at (e : trivialization F (π E)) [e.is_linear R] (b : B) : E b →ₗ[R] F := +protected def linear_map_at (e : trivialization F (π F E)) [e.is_linear R] (b : B) : E b →ₗ[R] F := e.to_pretrivialization.linear_map_at R b variables {R} -lemma coe_linear_map_at (e : trivialization F (π E)) [e.is_linear R] (b : B) : - ⇑(e.linear_map_at R b) = λ y, if b ∈ e.base_set then (e (total_space_mk b y)).2 else 0 := +lemma coe_linear_map_at (e : trivialization F (π F E)) [e.is_linear R] (b : B) : + ⇑(e.linear_map_at R b) = λ y, if b ∈ e.base_set then (e ⟨b, y⟩).2 else 0 := e.to_pretrivialization.coe_linear_map_at b -lemma coe_linear_map_at_of_mem (e : trivialization F (π E)) [e.is_linear R] {b : B} +lemma coe_linear_map_at_of_mem (e : trivialization F (π F E)) [e.is_linear R] {b : B} (hb : b ∈ e.base_set) : - ⇑(e.linear_map_at R b) = λ y, (e (total_space_mk b y)).2 := + ⇑(e.linear_map_at R b) = λ y, (e ⟨b, y⟩).2 := by simp_rw [coe_linear_map_at, if_pos hb] -lemma linear_map_at_apply (e : trivialization F (π E)) [e.is_linear R] {b : B} (y : E b) : - e.linear_map_at R b y = if b ∈ e.base_set then (e (total_space_mk b y)).2 else 0 := +lemma linear_map_at_apply (e : trivialization F (π F E)) [e.is_linear R] {b : B} (y : E b) : + e.linear_map_at R b y = if b ∈ e.base_set then (e ⟨b, y⟩).2 else 0 := by rw [coe_linear_map_at] -lemma linear_map_at_def_of_mem (e : trivialization F (π E)) [e.is_linear R] {b : B} +lemma linear_map_at_def_of_mem (e : trivialization F (π F E)) [e.is_linear R] {b : B} (hb : b ∈ e.base_set) : e.linear_map_at R b = e.linear_equiv_at R b hb := dif_pos hb -lemma linear_map_at_def_of_not_mem (e : trivialization F (π E)) [e.is_linear R] {b : B} +lemma linear_map_at_def_of_not_mem (e : trivialization F (π F E)) [e.is_linear R] {b : B} (hb : b ∉ e.base_set) : e.linear_map_at R b = 0 := dif_neg hb -lemma symmₗ_linear_map_at (e : trivialization F (π E)) [e.is_linear R] {b : B} (hb : b ∈ e.base_set) - (y : E b) : +lemma symmₗ_linear_map_at (e : trivialization F (π F E)) [e.is_linear R] {b : B} + (hb : b ∈ e.base_set) (y : E b) : e.symmₗ R b (e.linear_map_at R b y) = y := e.to_pretrivialization.symmₗ_linear_map_at hb y -lemma linear_map_at_symmₗ (e : trivialization F (π E)) [e.is_linear R] {b : B} (hb : b ∈ e.base_set) - (y : F) : +lemma linear_map_at_symmₗ (e : trivialization F (π F E)) [e.is_linear R] {b : B} + (hb : b ∈ e.base_set) (y : F) : e.linear_map_at R b (e.symmₗ R b y) = y := e.to_pretrivialization.linear_map_at_symmₗ hb y @@ -252,7 +254,7 @@ variables (R) /-- A coordinate change function between two trivializations, as a continuous linear equivalence. Defined to be the identity when `b` does not lie in the base set of both trivializations. -/ -def coord_changeL (e e' : trivialization F (π E)) [e.is_linear R] [e'.is_linear R] (b : B) : +def coord_changeL (e e' : trivialization F (π F E)) [e.is_linear R] [e'.is_linear R] (b : B) : F ≃L[R] F := { continuous_to_fun := begin by_cases hb : b ∈ e.base_set ∩ e'.base_set, @@ -278,19 +280,19 @@ def coord_changeL (e e' : trivialization F (π E)) [e.is_linear R] [e'.is_linear variables {R} -lemma coe_coord_changeL (e e' : trivialization F (π E)) [e.is_linear R] [e'.is_linear R] {b : B} +lemma coe_coord_changeL (e e' : trivialization F (π F E)) [e.is_linear R] [e'.is_linear R] {b : B} (hb : b ∈ e.base_set ∩ e'.base_set) : ⇑(coord_changeL R e e' b) = (e.linear_equiv_at R b hb.1).symm.trans (e'.linear_equiv_at R b hb.2) := congr_arg linear_equiv.to_fun (dif_pos hb) -lemma coe_coord_changeL' (e e' : trivialization F (π E)) [e.is_linear R] [e'.is_linear R] {b : B} +lemma coe_coord_changeL' (e e' : trivialization F (π F E)) [e.is_linear R] [e'.is_linear R] {b : B} (hb : b ∈ e.base_set ∩ e'.base_set) : (coord_changeL R e e' b).to_linear_equiv = (e.linear_equiv_at R b hb.1).symm.trans (e'.linear_equiv_at R b hb.2) := linear_equiv.coe_injective (coe_coord_changeL _ _ _) -lemma symm_coord_changeL (e e' : trivialization F (π E)) [e.is_linear R] [e'.is_linear R] {b : B} +lemma symm_coord_changeL (e e' : trivialization F (π F E)) [e.is_linear R] [e'.is_linear R] {b : B} (hb : b ∈ e'.base_set ∩ e.base_set) : (e.coord_changeL R e' b).symm = e'.coord_changeL R e b := begin @@ -299,14 +301,14 @@ begin coe_coord_changeL' e e' hb.symm, linear_equiv.trans_symm, linear_equiv.symm_symm], end -lemma coord_changeL_apply (e e' : trivialization F (π E)) [e.is_linear R] [e'.is_linear R] {b : B} +lemma coord_changeL_apply (e e' : trivialization F (π F E)) [e.is_linear R] [e'.is_linear R] {b : B} (hb : b ∈ e.base_set ∩ e'.base_set) (y : F) : - coord_changeL R e e' b y = (e' (total_space_mk b (e.symm b y))).2 := + coord_changeL R e e' b y = (e' ⟨b, e.symm b y⟩).2 := congr_arg (λ f, linear_equiv.to_fun f y) (dif_pos hb) -lemma mk_coord_changeL (e e' : trivialization F (π E)) [e.is_linear R] [e'.is_linear R] {b : B} +lemma mk_coord_changeL (e e' : trivialization F (π F E)) [e.is_linear R] [e'.is_linear R] {b : B} (hb : b ∈ e.base_set ∩ e'.base_set) (y : F) : - (b, coord_changeL R e e' b y) = e' (total_space_mk b (e.symm b y)) := + (b, coord_changeL R e e' b y) = e' ⟨b, e.symm b y⟩ := begin ext, { rw [e.mk_symm hb.1 y, e'.coe_fst', e.proj_symm_apply' hb.1], @@ -314,19 +316,19 @@ begin { exact e.coord_changeL_apply e' hb y } end -lemma apply_symm_apply_eq_coord_changeL (e e' : trivialization F (π E)) [e.is_linear R] +lemma apply_symm_apply_eq_coord_changeL (e e' : trivialization F (π F E)) [e.is_linear R] [e'.is_linear R] {b : B} (hb : b ∈ e.base_set ∩ e'.base_set) (v : F) : e' (e.to_local_homeomorph.symm (b, v)) = (b, e.coord_changeL R e' b v) := by rw [e.mk_coord_changeL e' hb, e.mk_symm hb.1] /-- A version of `coord_change_apply` that fully unfolds `coord_change`. The right-hand side is ugly, but has good definitional properties for specifically defined trivializations. -/ -lemma coord_changeL_apply' (e e' : trivialization F (π E)) [e.is_linear R] [e'.is_linear R] {b : B} - (hb : b ∈ e.base_set ∩ e'.base_set) (y : F) : +lemma coord_changeL_apply' (e e' : trivialization F (π F E)) [e.is_linear R] [e'.is_linear R] + {b : B} (hb : b ∈ e.base_set ∩ e'.base_set) (y : F) : coord_changeL R e e' b y = (e' (e.to_local_homeomorph.symm (b, y))).2 := by rw [e.coord_changeL_apply e' hb, e.mk_symm hb.1] -lemma coord_changeL_symm_apply (e e' : trivialization F (π E)) [e.is_linear R] [e'.is_linear R] +lemma coord_changeL_symm_apply (e e' : trivialization F (π F E)) [e.is_linear R] [e'.is_linear R] {b : B} (hb : b ∈ e.base_set ∩ e'.base_set) : ⇑(coord_changeL R e e' b).symm = (e'.linear_equiv_at R b hb.2).symm.trans (e.linear_equiv_at R b hb.1) := @@ -341,29 +343,29 @@ section namespace bundle /-- The zero section of a vector bundle -/ -def zero_section [∀ x, has_zero (E x)] : B → total_space E := -λ x, total_space_mk x 0 +def zero_section [∀ x, has_zero (E x)] : B → total_space F E := +λ x, ⟨x, 0⟩ @[simp, mfld_simps] -lemma zero_section_proj [∀ x, has_zero (E x)] (x : B) : (zero_section E x).proj = x := rfl +lemma zero_section_proj [∀ x, has_zero (E x)] (x : B) : (zero_section F E x).proj = x := rfl @[simp, mfld_simps] -lemma zero_section_snd [∀ x, has_zero (E x)] (x : B) : (zero_section E x).2 = 0 := rfl +lemma zero_section_snd [∀ x, has_zero (E x)] (x : B) : (zero_section F E x).2 = 0 := rfl end bundle open bundle variables [nontrivially_normed_field R] [∀ x, add_comm_monoid (E x)] [∀ x, module R (E x)] [normed_add_comm_group F] [normed_space R F] [topological_space B] - [topological_space (total_space E)] [∀ x, topological_space (E x)] [fiber_bundle F E] + [topological_space (total_space F E)] [∀ x, topological_space (E x)] [fiber_bundle F E] -/-- The space `total_space E` (for `E : B → Type*` such that each `E x` is a topological vector +/-- The space `total_space F E` (for `E : B → Type*` such that each `E x` is a topological vector space) has a topological vector space structure with fiber `F` (denoted with `vector_bundle R F E`) if around every point there is a fiber bundle trivialization which is linear in the fibers. -/ class vector_bundle : Prop := -(trivialization_linear' : ∀ (e : trivialization F (π E)) [mem_trivialization_atlas e], +(trivialization_linear' : ∀ (e : trivialization F (π F E)) [mem_trivialization_atlas e], e.is_linear R) -(continuous_on_coord_change' [] : ∀ (e e' : trivialization F (π E)) [mem_trivialization_atlas e] +(continuous_on_coord_change' [] : ∀ (e e' : trivialization F (π F E)) [mem_trivialization_atlas e] [mem_trivialization_atlas e'], continuous_on (λ b, by exactI trivialization.coord_changeL R e e' b : B → F →L[R] F) (e.base_set ∩ e'.base_set)) @@ -371,12 +373,12 @@ class vector_bundle : Prop := variables {F E} @[priority 100] -instance trivialization_linear [vector_bundle R F E] (e : trivialization F (π E)) +instance trivialization_linear [vector_bundle R F E] (e : trivialization F (π F E)) [mem_trivialization_atlas e] : e.is_linear R := vector_bundle.trivialization_linear' e -lemma continuous_on_coord_change [vector_bundle R F E] (e e' : trivialization F (π E)) +lemma continuous_on_coord_change [vector_bundle R F E] (e e' : trivialization F (π F E)) [he : mem_trivialization_atlas e] [he' : mem_trivialization_atlas e'] : continuous_on @@ -388,7 +390,7 @@ namespace trivialization /-- Forward map of `continuous_linear_equiv_at` (only propositionally equal), defined everywhere (`0` outside domain). -/ @[simps apply {fully_applied := ff}] -def continuous_linear_map_at (e : trivialization F (π E)) [e.is_linear R] (b : B) : +def continuous_linear_map_at (e : trivialization F (π F E)) [e.is_linear R] (b : B) : E b →L[R] F := { to_fun := e.linear_map_at R b, -- given explicitly to help `simps` cont := begin @@ -403,7 +405,7 @@ def continuous_linear_map_at (e : trivialization F (π E)) [e.is_linear R] (b : /-- Backwards map of `continuous_linear_equiv_at`, defined everywhere. -/ @[simps apply {fully_applied := ff}] -def symmL (e : trivialization F (π E)) [e.is_linear R] (b : B) : F →L[R] E b := +def symmL (e : trivialization F (π F E)) [e.is_linear R] (b : B) : F →L[R] E b := { to_fun := e.symm b, -- given explicitly to help `simps` cont := begin by_cases hb : b ∈ e.base_set, @@ -416,12 +418,12 @@ def symmL (e : trivialization F (π E)) [e.is_linear R] (b : B) : F →L[R] E b variables {R} -lemma symmL_continuous_linear_map_at (e : trivialization F (π E)) [e.is_linear R] {b : B} +lemma symmL_continuous_linear_map_at (e : trivialization F (π F E)) [e.is_linear R] {b : B} (hb : b ∈ e.base_set) (y : E b) : e.symmL R b (e.continuous_linear_map_at R b y) = y := e.symmₗ_linear_map_at hb y -lemma continuous_linear_map_at_symmL (e : trivialization F (π E)) [e.is_linear R] {b : B} +lemma continuous_linear_map_at_symmL (e : trivialization F (π F E)) [e.is_linear R] {b : B} (hb : b ∈ e.base_set) (y : F) : e.continuous_linear_map_at R b (e.symmL R b y) = y := e.linear_map_at_symmₗ hb y @@ -431,9 +433,9 @@ variables (R) /-- In a vector bundle, a trivialization in the fiber (which is a priori only linear) is in fact a continuous linear equiv between the fibers and the model fiber. -/ @[simps apply symm_apply {fully_applied := ff}] -def continuous_linear_equiv_at (e : trivialization F (π E)) [e.is_linear R] (b : B) +def continuous_linear_equiv_at (e : trivialization F (π F E)) [e.is_linear R] (b : B) (hb : b ∈ e.base_set) : E b ≃L[R] F := -{ to_fun := λ y, (e (total_space_mk b y)).2, -- given explicitly to help `simps` +{ to_fun := λ y, (e ⟨b, y⟩).2, -- given explicitly to help `simps` inv_fun := e.symm b, -- given explicitly to help `simps` continuous_to_fun := continuous_snd.comp (e.continuous_on.comp_continuous (fiber_bundle.total_space_mk_inducing F E b).continuous @@ -443,24 +445,24 @@ def continuous_linear_equiv_at (e : trivialization F (π E)) [e.is_linear R] (b variables {R} -lemma coe_continuous_linear_equiv_at_eq (e : trivialization F (π E)) [e.is_linear R] {b : B} +lemma coe_continuous_linear_equiv_at_eq (e : trivialization F (π F E)) [e.is_linear R] {b : B} (hb : b ∈ e.base_set) : (e.continuous_linear_equiv_at R b hb : E b → F) = e.continuous_linear_map_at R b := (e.coe_linear_map_at_of_mem hb).symm -lemma symm_continuous_linear_equiv_at_eq (e : trivialization F (π E)) [e.is_linear R] {b : B} +lemma symm_continuous_linear_equiv_at_eq (e : trivialization F (π F E)) [e.is_linear R] {b : B} (hb : b ∈ e.base_set) : ((e.continuous_linear_equiv_at R b hb).symm : F → E b) = e.symmL R b := rfl -@[simp] lemma continuous_linear_equiv_at_apply' (e : trivialization F (π E)) [e.is_linear R] - (x : total_space E) (hx : x ∈ e.source) : +@[simp] lemma continuous_linear_equiv_at_apply' (e : trivialization F (π F E)) [e.is_linear R] + (x : total_space F E) (hx : x ∈ e.source) : e.continuous_linear_equiv_at R x.proj (e.mem_source.1 hx) x.2 = (e x).2 := by { cases x, refl } variables (R) -lemma apply_eq_prod_continuous_linear_equiv_at (e : trivialization F (π E)) [e.is_linear R] (b : B) - (hb : b ∈ e.base_set) (z : E b) : +lemma apply_eq_prod_continuous_linear_equiv_at (e : trivialization F (π F E)) [e.is_linear R] + (b : B) (hb : b ∈ e.base_set) (z : E b) : e ⟨b, z⟩ = (b, e.continuous_linear_equiv_at R b hb z) := begin ext, @@ -470,17 +472,17 @@ begin { simp only [coe_coe, continuous_linear_equiv_at_apply] } end -protected lemma zero_section (e : trivialization F (π E)) [e.is_linear R] - {x : B} (hx : x ∈ e.base_set) : e (zero_section E x) = (x, 0) := -by simp_rw [zero_section, total_space_mk, e.apply_eq_prod_continuous_linear_equiv_at R x hx 0, +protected lemma zero_section (e : trivialization F (π F E)) [e.is_linear R] + {x : B} (hx : x ∈ e.base_set) : e (zero_section F E x) = (x, 0) := +by simp_rw [zero_section, e.apply_eq_prod_continuous_linear_equiv_at R x hx 0, map_zero] variables {R} -lemma symm_apply_eq_mk_continuous_linear_equiv_at_symm (e : trivialization F (π E)) [e.is_linear R] - (b : B) (hb : b ∈ e.base_set) (z : F) : +lemma symm_apply_eq_mk_continuous_linear_equiv_at_symm (e : trivialization F (π F E)) + [e.is_linear R] (b : B) (hb : b ∈ e.base_set) (z : F) : e.to_local_homeomorph.symm ⟨b, z⟩ - = total_space_mk b ((e.continuous_linear_equiv_at R b hb).symm z) := + = ⟨b, (e.continuous_linear_equiv_at R b hb).symm z⟩ := begin have h : (b, z) ∈ e.target, { rw e.target_eq, @@ -491,7 +493,7 @@ begin continuous_linear_equiv.apply_symm_apply], end -lemma comp_continuous_linear_equiv_at_eq_coord_change (e e' : trivialization F (π E)) +lemma comp_continuous_linear_equiv_at_eq_coord_change (e e' : trivialization F (π F E)) [e.is_linear R] [e'.is_linear R] {b : B} (hb : b ∈ e.base_set ∩ e'.base_set) : (e.continuous_linear_equiv_at R b hb.1).symm.trans (e'.continuous_linear_equiv_at R b hb.2) = coord_changeL R e e' b := @@ -578,13 +580,12 @@ instance add_comm_group_fiber [add_comm_group F] : ∀ (x : B), add_comm_group ( by dsimp [vector_bundle_core.fiber]; delta_instance fiber_bundle_core.fiber /-- The projection from the total space of a fiber bundle core, on its base. -/ -@[reducible, simp, mfld_simps] protected def proj : total_space Z.fiber → B := total_space.proj +@[reducible, simp, mfld_simps] protected def proj : total_space F Z.fiber → B := total_space.proj /-- The total space of the vector bundle, as a convenience function for dot notation. -It is by definition equal to `bundle.total_space Z.fiber`, a.k.a. `Σ x, Z.fiber x` but with a -different name for typeclass inference. -/ +It is by definition equal to `bundle.total_space Z.fiber`. -/ @[nolint unused_arguments, reducible] -protected def total_space := bundle.total_space Z.fiber +protected def total_space := bundle.total_space F Z.fiber /-- Local homeomorphism version of the trivialization change. -/ def triv_change (i j : ι) : local_homeomorph (B × F) (B × F) := @@ -606,7 +607,7 @@ variables (b : B) (a : F) /-- One of the standard local trivializations of a vector bundle constructed from core, taken by considering this in particular as a fiber bundle constructed from core. -/ -def local_triv (i : ι) : trivialization F (π Z.fiber) := +def local_triv (i : ι) : trivialization F (π F Z.fiber) := by dsimp [vector_bundle_core.total_space, vector_bundle_core.fiber]; exact Z.to_fiber_bundle_core.local_triv i @@ -650,7 +651,7 @@ end /-- Preferred local trivialization of a vector bundle constructed from core, at a given point, as a bundle trivialization -/ -def local_triv_at (b : B) : trivialization F (π Z.fiber) := +def local_triv_at (b : B) : trivialization F (π F Z.fiber) := Z.local_triv (Z.index_at b) @[simp, mfld_simps] lemma local_triv_at_def : @@ -750,17 +751,17 @@ This makes it inconvenient to explicitly define a `coord_change` function when c `vector_prebundle`. -/ @[nolint has_nonempty_instance] structure vector_prebundle := -(pretrivialization_atlas : set (pretrivialization F (π E))) -(pretrivialization_linear' : ∀ (e : pretrivialization F (π E)) (he : e ∈ pretrivialization_atlas), +(pretrivialization_atlas : set (pretrivialization F (π F E))) +(pretrivialization_linear' : ∀ (e : pretrivialization F (π F E)) (he : e ∈ pretrivialization_atlas), e.is_linear R) -(pretrivialization_at : B → pretrivialization F (π E)) +(pretrivialization_at : B → pretrivialization F (π F E)) (mem_base_pretrivialization_at : ∀ x : B, x ∈ (pretrivialization_at x).base_set) (pretrivialization_mem_atlas : ∀ x : B, pretrivialization_at x ∈ pretrivialization_atlas) (exists_coord_change : ∀ (e e' ∈ pretrivialization_atlas), ∃ f : B → F →L[R] F, continuous_on f (e.base_set ∩ e'.base_set) ∧ ∀ (b : B) (hb : b ∈ e.base_set ∩ e'.base_set) (v : F), - f b v = (e' (total_space_mk b (e.symm b v))).2) -(total_space_mk_inducing : ∀ (b : B), inducing ((pretrivialization_at b) ∘ (total_space_mk b))) + f b v = (e' ⟨b, e.symm b v⟩).2) +(total_space_mk_inducing : ∀ (b : B), inducing ((pretrivialization_at b) ∘ (total_space.mk b))) namespace vector_prebundle @@ -769,26 +770,26 @@ variables {R E F} /-- A randomly chosen coordinate change on a `vector_prebundle`, given by the field `exists_coord_change`. -/ def coord_change (a : vector_prebundle R F E) - {e e' : pretrivialization F (π E)} (he : e ∈ a.pretrivialization_atlas) + {e e' : pretrivialization F (π F E)} (he : e ∈ a.pretrivialization_atlas) (he' : e' ∈ a.pretrivialization_atlas) (b : B) : F →L[R] F := classical.some (a.exists_coord_change e he e' he') b lemma continuous_on_coord_change (a : vector_prebundle R F E) - {e e' : pretrivialization F (π E)} (he : e ∈ a.pretrivialization_atlas) + {e e' : pretrivialization F (π F E)} (he : e ∈ a.pretrivialization_atlas) (he' : e' ∈ a.pretrivialization_atlas) : continuous_on (a.coord_change he he') (e.base_set ∩ e'.base_set) := (classical.some_spec (a.exists_coord_change e he e' he')).1 lemma coord_change_apply (a : vector_prebundle R F E) - {e e' : pretrivialization F (π E)} (he : e ∈ a.pretrivialization_atlas) + {e e' : pretrivialization F (π F E)} (he : e ∈ a.pretrivialization_atlas) (he' : e' ∈ a.pretrivialization_atlas) {b : B} (hb : b ∈ e.base_set ∩ e'.base_set) (v : F) : - a.coord_change he he' b v = (e' (total_space_mk b (e.symm b v))).2 := + a.coord_change he he' b v = (e' ⟨b, e.symm b v⟩).2 := (classical.some_spec (a.exists_coord_change e he e' he')).2 b hb v lemma mk_coord_change (a : vector_prebundle R F E) - {e e' : pretrivialization F (π E)} (he : e ∈ a.pretrivialization_atlas) + {e e' : pretrivialization F (π F E)} (he : e ∈ a.pretrivialization_atlas) (he' : e' ∈ a.pretrivialization_atlas) {b : B} (hb : b ∈ e.base_set ∩ e'.base_set) (v : F) : - (b, a.coord_change he he' b v) = e' (total_space_mk b (e.symm b v)) := + (b, a.coord_change he he' b v) = e' ⟨b, e.symm b v⟩ := begin ext, { rw [e.mk_symm hb.1 v, e'.coe_fst', e.proj_symm_apply' hb.1], @@ -820,18 +821,18 @@ def to_fiber_prebundle (a : vector_prebundle R F E) : /-- Topology on the total space that will make the prebundle into a bundle. -/ def total_space_topology (a : vector_prebundle R F E) : - topological_space (total_space E) := + topological_space (total_space F E) := a.to_fiber_prebundle.total_space_topology /-- Promotion from a `trivialization` in the `pretrivialization_atlas` of a `vector_prebundle` to a `trivialization`. -/ def trivialization_of_mem_pretrivialization_atlas (a : vector_prebundle R F E) - {e : pretrivialization F (π E)} (he : e ∈ a.pretrivialization_atlas) : - @trivialization B F _ _ _ a.total_space_topology (π E) := + {e : pretrivialization F (π F E)} (he : e ∈ a.pretrivialization_atlas) : + @trivialization B F _ _ _ a.total_space_topology (π F E) := a.to_fiber_prebundle.trivialization_of_mem_pretrivialization_atlas he lemma linear_of_mem_pretrivialization_atlas (a : vector_prebundle R F E) - {e : pretrivialization F (π E)} (he : e ∈ a.pretrivialization_atlas) : + {e : pretrivialization F (π F E)} (he : e ∈ a.pretrivialization_atlas) : @trivialization.is_linear R B F _ _ _ _ a.total_space_topology _ _ _ _ (trivialization_of_mem_pretrivialization_atlas a he) := { linear := (a.pretrivialization_linear' e he).linear } @@ -839,15 +840,15 @@ lemma linear_of_mem_pretrivialization_atlas (a : vector_prebundle R F E) variable (a : vector_prebundle R F E) lemma mem_trivialization_at_source (b : B) (x : E b) : - total_space_mk b x ∈ (a.pretrivialization_at b).source := + total_space.mk b x ∈ (a.pretrivialization_at b).source := a.to_fiber_prebundle.mem_trivialization_at_source b x @[simp] lemma total_space_mk_preimage_source (b : B) : - (total_space_mk b) ⁻¹' (a.pretrivialization_at b).source = univ := + (total_space.mk b) ⁻¹' (a.pretrivialization_at b).source = univ := a.to_fiber_prebundle.total_space_mk_preimage_source b @[continuity] lemma continuous_total_space_mk (b : B) : - @continuous _ _ _ a.total_space_topology (total_space_mk b) := + @continuous _ _ _ a.total_space_topology (total_space.mk b) := a.to_fiber_prebundle.continuous_total_space_mk b /-- Make a `fiber_bundle` from a `vector_prebundle`; auxiliary construction for @@ -884,10 +885,10 @@ variables {𝕜₁ 𝕜₂ : Type*} [nontrivially_normed_field 𝕜₁] [nontriv variables {σ : 𝕜₁ →+* 𝕜₂} variables {B' : Type*} [topological_space B'] -variables [normed_space 𝕜₁ F] [Π x, module 𝕜₁ (E x)] [topological_space (total_space E)] +variables [normed_space 𝕜₁ F] [Π x, module 𝕜₁ (E x)] [topological_space (total_space F E)] variables {F' : Type*} [normed_add_comm_group F'] [normed_space 𝕜₂ F'] {E' : B' → Type*} [Π x, add_comm_monoid (E' x)] [Π x, module 𝕜₂ (E' x)] - [topological_space (total_space E')] + [topological_space (total_space F' E')] variables [fiber_bundle F E] [vector_bundle 𝕜₁ F E] variables [Π x, topological_space (E' x)] [fiber_bundle F' E'] [vector_bundle 𝕜₂ F' E'] variables (F E F' E') diff --git a/src/topology/vector_bundle/constructions.lean b/src/topology/vector_bundle/constructions.lean index 08067b91091e8..210ae668d9ec6 100644 --- a/src/topology/vector_bundle/constructions.lean +++ b/src/topology/vector_bundle/constructions.lean @@ -75,15 +75,15 @@ end bundle.trivial section variables (𝕜 : Type*) {B : Type*} [nontrivially_normed_field 𝕜] [topological_space B] (F₁ : Type*) [normed_add_comm_group F₁] [normed_space 𝕜 F₁] - (E₁ : B → Type*) [topological_space (total_space E₁)] + (E₁ : B → Type*) [topological_space (total_space F₁ E₁)] (F₂ : Type*) [normed_add_comm_group F₂] [normed_space 𝕜 F₂] - (E₂ : B → Type*) [topological_space (total_space E₂)] + (E₂ : B → Type*) [topological_space (total_space F₂ E₂)] namespace trivialization variables {F₁ E₁ F₂ E₂} [Π x, add_comm_monoid (E₁ x)] [Π x, module 𝕜 (E₁ x)] [Π x, add_comm_monoid (E₂ x)] [Π x, module 𝕜 (E₂ x)] - (e₁ e₁' : trivialization F₁ (π E₁)) (e₂ e₂' : trivialization F₂ (π E₂)) + (e₁ e₁' : trivialization F₁ (π F₁ E₁)) (e₂ e₂' : trivialization F₂ (π F₂ E₂)) instance prod.is_linear [e₁.is_linear 𝕜] [e₂.is_linear 𝕜] : (e₁.prod e₂).is_linear 𝕜 := { linear := λ x ⟨h₁, h₂⟩, (((e₁.linear 𝕜 h₁).mk' _).prod_map ((e₂.linear 𝕜 h₂).mk' _)).is_linear } @@ -146,9 +146,9 @@ instance vector_bundle.prod [vector_bundle 𝕜 F₁ E₁] [vector_bundle 𝕜 variables {𝕜 F₁ E₁ F₂ E₂} -@[simp] lemma trivialization.continuous_linear_equiv_at_prod {e₁ : trivialization F₁ (π E₁)} - {e₂ : trivialization F₂ (π E₂)} [e₁.is_linear 𝕜] [e₂.is_linear 𝕜] {x : B} (hx₁ : x ∈ e₁.base_set) - (hx₂ : x ∈ e₂.base_set) : +@[simp] lemma trivialization.continuous_linear_equiv_at_prod {e₁ : trivialization F₁ (π F₁ E₁)} + {e₂ : trivialization F₂ (π F₂ E₂)} [e₁.is_linear 𝕜] [e₂.is_linear 𝕜] {x : B} + (hx₁ : x ∈ e₁.base_set) (hx₂ : x ∈ e₂.base_set) : (e₁.prod e₂).continuous_linear_equiv_at 𝕜 x ⟨hx₁, hx₂⟩ = (e₁.continuous_linear_equiv_at 𝕜 x hx₁).prod (e₂.continuous_linear_equiv_at 𝕜 x hx₂) := begin @@ -172,12 +172,12 @@ instance [semiring R] [∀ (x : B), add_comm_monoid (E x)] [∀ x, module R (E x ∀ (x : B'), module R ((f *ᵖ E) x) := by delta_instance bundle.pullback -variables {E F} [topological_space B'] [topological_space (total_space E)] +variables {E F} [topological_space B'] [topological_space (total_space F E)] [nontrivially_normed_field 𝕜] [normed_add_comm_group F] [normed_space 𝕜 F] [topological_space B] [∀ x, add_comm_monoid (E x)] [∀ x, module 𝕜 (E x)] {K : Type*} [continuous_map_class K B' B] -instance trivialization.pullback_linear (e : trivialization F (π E)) [e.is_linear 𝕜] (f : K) : +instance trivialization.pullback_linear (e : trivialization F (π F E)) [e.is_linear 𝕜] (f : K) : (@trivialization.pullback _ _ _ B' _ _ _ _ _ _ _ e f).is_linear 𝕜 := { linear := λ x h, e.linear 𝕜 h } diff --git a/src/topology/vector_bundle/hom.lean b/src/topology/vector_bundle/hom.lean index 68258be0aa117..7e09babbf0ef9 100644 --- a/src/topology/vector_bundle/hom.lean +++ b/src/topology/vector_bundle/hom.lean @@ -40,62 +40,39 @@ noncomputable theory open_locale bundle open bundle set continuous_linear_map -section defs -variables {𝕜₁ 𝕜₂ : Type*} [normed_field 𝕜₁] [normed_field 𝕜₂] -variables (σ : 𝕜₁ →+* 𝕜₂) -variables {B : Type*} -variables (F₁ : Type*) (E₁ : B → Type*) [Π x, add_comm_group (E₁ x)] [Π x, module 𝕜₁ (E₁ x)] -variables [Π x, topological_space (E₁ x)] -variables (F₂ : Type*) (E₂ : B → Type*) [Π x, add_comm_group (E₂ x)] [Π x, module 𝕜₂ (E₂ x)] -variables [Π x, topological_space (E₂ x)] - -include F₁ F₂ - --- In this definition we require the scalar rings `𝕜₁` and `𝕜₂` to be normed fields, although --- something much weaker (maybe `comm_semiring`) would suffice mathematically -- this is because of --- a typeclass inference bug with pi-types: --- https://leanprover.zulipchat.com/#narrow/stream/116395-maths/topic/vector.20bundles.20--.20typeclass.20inference.20issue -/-- The bundle of continuous `σ`-semilinear maps between the topological vector bundles `E₁` and -`E₂`. This is a type synonym for `λ x, E₁ x →SL[σ] E₂ x`. - -We intentionally add `F₁` and `F₂` as arguments to this type, so that instances on this type -(that depend on `F₁` and `F₂`) actually refer to `F₁` and `F₂`. -/ -@[derive inhabited, nolint unused_arguments] -protected def bundle.continuous_linear_map (x : B) : Type* := -E₁ x →SL[σ] E₂ x - -instance bundle.continuous_linear_map.add_monoid_hom_class (x : B) : - add_monoid_hom_class (bundle.continuous_linear_map σ F₁ E₁ F₂ E₂ x) (E₁ x) (E₂ x) := -by delta_instance bundle.continuous_linear_map - -variables [Π x, topological_add_group (E₂ x)] - -instance (x : B) : topological_space (bundle.continuous_linear_map σ F₁ E₁ F₂ E₂ x) := -by delta_instance bundle.continuous_linear_map - -instance (x : B) : add_comm_monoid (bundle.continuous_linear_map σ F₁ E₁ F₂ E₂ x) := -by delta_instance bundle.continuous_linear_map - -variables [∀ x, has_continuous_smul 𝕜₂ (E₂ x)] - -instance (x : B) : module 𝕜₂ (bundle.continuous_linear_map σ F₁ E₁ F₂ E₂ x) := -by delta_instance bundle.continuous_linear_map - -end defs - variables {𝕜₁ : Type*} [nontrivially_normed_field 𝕜₁] {𝕜₂ : Type*} [nontrivially_normed_field 𝕜₂] (σ : 𝕜₁ →+* 𝕜₂) [iσ : ring_hom_isometric σ] -variables {B : Type*} [topological_space B] +variables {B : Type*} -variables (F₁ : Type*) [normed_add_comm_group F₁] [normed_space 𝕜₁ F₁] +variables {F₁ : Type*} [normed_add_comm_group F₁] [normed_space 𝕜₁ F₁] (E₁ : B → Type*) [Π x, add_comm_group (E₁ x)] [Π x, module 𝕜₁ (E₁ x)] - [topological_space (total_space E₁)] -variables (F₂ : Type*) [normed_add_comm_group F₂][normed_space 𝕜₂ F₂] + [topological_space (total_space F₁ E₁)] +variables {F₂ : Type*} [normed_add_comm_group F₂] [normed_space 𝕜₂ F₂] (E₂ : B → Type*) [Π x, add_comm_group (E₂ x)] [Π x, module 𝕜₂ (E₂ x)] - [topological_space (total_space E₂)] + [topological_space (total_space F₂ E₂)] + +/-- A reducible type synonym for the bundle of continuous (semi)linear maps. For some reason, it +helps with instance search. + +Porting note: after the port is done, we may want to remove this definition. +-/ +@[reducible] +protected def bundle.continuous_linear_map [∀ x, topological_space (E₁ x)] + [∀ x, topological_space (E₂ x)] : Π x : B, Type* := +λ x, E₁ x →SL[σ] E₂ x + +-- Porting note: possibly remove after the port +instance bundle.continuous_linear_map.module [∀ x, topological_space (E₁ x)] + [∀ x, topological_space (E₂ x)] [∀ x, topological_add_group (E₂ x)] + [∀ x, has_continuous_const_smul 𝕜₂ (E₂ x)] : + ∀ x, module 𝕜₂ (bundle.continuous_linear_map σ E₁ E₂ x) := +λ _, infer_instance + +variables {E₁ E₂} -variables {F₁ E₁ F₂ E₂} (e₁ e₁' : trivialization F₁ (π E₁)) (e₂ e₂' : trivialization F₂ (π E₂)) +variables [topological_space B] (e₁ e₁' : trivialization F₁ (π F₁ E₁)) + (e₂ e₂' : trivialization F₂ (π F₂ E₂)) namespace pretrivialization @@ -149,7 +126,7 @@ continuous `σ`-semilinear maps from `E₁` to `E₂`. That is, the map which wi trivialization, after the bundle of continuous semilinear maps is equipped with the right topological vector bundle structure. -/ def continuous_linear_map : - pretrivialization (F₁ →SL[σ] F₂) (π (bundle.continuous_linear_map σ F₁ E₁ F₂ E₂)) := + pretrivialization (F₁ →SL[σ] F₂) (π (F₁ →SL[σ] F₂) (bundle.continuous_linear_map σ E₁ E₂)) := { to_fun := λ p, ⟨p.1, continuous_linear_map.comp (e₂.continuous_linear_map_at 𝕜₂ p.1) (p.2.comp (e₁.symmL 𝕜₁ p.1 : F₁ →L[𝕜₁] E₁ p.1) : F₁ →SL[σ] E₂ p.1)⟩, inv_fun := λ p, ⟨p.1, continuous_linear_map.comp (e₂.symmL 𝕜₂ p.1) @@ -179,6 +156,7 @@ def continuous_linear_map : include ita +-- porting note: todo: see if Lean 4 can generate this instance without a hint instance continuous_linear_map.is_linear [Π x, has_continuous_add (E₂ x)] [Π x, has_continuous_smul 𝕜₂ (E₂ x)] : (pretrivialization.continuous_linear_map σ e₁ e₂).is_linear 𝕜₂ := @@ -199,7 +177,7 @@ instance continuous_linear_map.is_linear omit ita lemma continuous_linear_map_apply - (p : total_space (bundle.continuous_linear_map σ F₁ E₁ F₂ E₂)) : + (p : total_space (F₁ →SL[σ] F₂) (λ x, E₁ x →SL[σ] E₂ x)) : (continuous_linear_map σ e₁ e₂) p = ⟨p.1, continuous_linear_map.comp (e₂.continuous_linear_map_at 𝕜₂ p.1) (p.2.comp (e₁.symmL 𝕜₁ p.1 : F₁ →L[𝕜₁] E₁ p.1) : F₁ →SL[σ] E₂ p.1)⟩ := @@ -224,8 +202,7 @@ end lemma continuous_linear_map_coord_change_apply (b : B) (hb : b ∈ (e₁.base_set ∩ e₂.base_set) ∩ (e₁'.base_set ∩ e₂'.base_set)) (L : F₁ →SL[σ] F₂) : continuous_linear_map_coord_change σ e₁ e₁' e₂ e₂' b L = - (continuous_linear_map σ e₁' e₂' - (total_space_mk b ((continuous_linear_map σ e₁ e₂).symm b L))).2 := + (continuous_linear_map σ e₁' e₂' ⟨b, ((continuous_linear_map σ e₁ e₂).symm b L)⟩).2 := begin ext v, simp_rw [continuous_linear_map_coord_change, continuous_linear_equiv.coe_coe, @@ -234,7 +211,6 @@ begin continuous_linear_map_apply, continuous_linear_map_symm_apply' σ e₁ e₂ hb.1, comp_apply, continuous_linear_equiv.coe_coe, continuous_linear_equiv.symm_symm, trivialization.continuous_linear_map_at_apply, trivialization.symmL_apply], - dsimp only [total_space_mk], rw [e₂.coord_changeL_apply e₂', e₁'.coord_changeL_apply e₁, e₁.coe_linear_map_at_of_mem hb.1.1, e₂'.coe_linear_map_at_of_mem hb.2.2], exacts [⟨hb.2.1, hb.1.1⟩, ⟨hb.1.2, hb.2.2⟩] @@ -255,10 +231,9 @@ include iσ `vector_bundle` instance, in which the pretrivializations are collated but no topology on the total space is yet provided). -/ def _root_.bundle.continuous_linear_map.vector_prebundle : - vector_prebundle 𝕜₂ (F₁ →SL[σ] F₂) - (bundle.continuous_linear_map σ F₁ E₁ F₂ E₂) := + vector_prebundle 𝕜₂ (F₁ →SL[σ] F₂) (bundle.continuous_linear_map σ E₁ E₂) := { pretrivialization_atlas := - {e | ∃ (e₁ : trivialization F₁ (π E₁)) (e₂ : trivialization F₂ (π E₂)) + {e | ∃ (e₁ : trivialization F₁ (π F₁ E₁)) (e₂ : trivialization F₂ (π F₂ E₂)) [mem_trivialization_atlas e₁] [mem_trivialization_atlas e₂], by exactI e = pretrivialization.continuous_linear_map σ e₁ e₂}, pretrivialization_linear' := begin @@ -279,7 +254,6 @@ def _root_.bundle.continuous_linear_map.vector_prebundle : total_space_mk_inducing := begin intros b, - dsimp [bundle.continuous_linear_map.topological_space, bundle.continuous_linear_map], let L₁ : E₁ b ≃L[𝕜₁] F₁ := (trivialization_at F₁ E₁ b).continuous_linear_equiv_at 𝕜₁ b (mem_base_set_trivialization_at _ _ _), let L₂ : E₂ b ≃L[𝕜₂] F₂ := (trivialization_at F₂ E₂ b).continuous_linear_equiv_at 𝕜₂ b @@ -298,19 +272,19 @@ def _root_.bundle.continuous_linear_map.vector_prebundle : /-- Topology on the total space of the continuous `σ`-semilinear_maps between two "normable" vector bundles over the same base. -/ instance bundle.continuous_linear_map.topological_space_total_space : - topological_space (total_space (bundle.continuous_linear_map σ F₁ E₁ F₂ E₂)) := + topological_space (total_space (F₁ →SL[σ] F₂) (bundle.continuous_linear_map σ E₁ E₂)) := (bundle.continuous_linear_map.vector_prebundle σ F₁ E₁ F₂ E₂).total_space_topology /-- The continuous `σ`-semilinear_maps between two vector bundles form a fiber bundle. -/ instance _root_.bundle.continuous_linear_map.fiber_bundle : - fiber_bundle (F₁ →SL[σ] F₂) (bundle.continuous_linear_map σ F₁ E₁ F₂ E₂) := + fiber_bundle (F₁ →SL[σ] F₂) (λ x, E₁ x →SL[σ] E₂ x) := (bundle.continuous_linear_map.vector_prebundle σ F₁ E₁ F₂ E₂).to_fiber_bundle /-- The continuous `σ`-semilinear_maps between two vector bundles form a vector bundle. -/ instance _root_.bundle.continuous_linear_map.vector_bundle : - vector_bundle 𝕜₂ (F₁ →SL[σ] F₂) (bundle.continuous_linear_map σ F₁ E₁ F₂ E₂) := + vector_bundle 𝕜₂ (F₁ →SL[σ] F₂) (bundle.continuous_linear_map σ E₁ E₂) := (bundle.continuous_linear_map.vector_prebundle σ F₁ E₁ F₂ E₂).to_vector_bundle @@ -323,12 +297,12 @@ include he₁ he₂ the induced trivialization for the continuous `σ`-semilinear maps from `E₁` to `E₂`, whose base set is `e₁.base_set ∩ e₂.base_set`. -/ def trivialization.continuous_linear_map : - trivialization (F₁ →SL[σ] F₂) (π (bundle.continuous_linear_map σ F₁ E₁ F₂ E₂)) := + trivialization (F₁ →SL[σ] F₂) (π (F₁ →SL[σ] F₂) (bundle.continuous_linear_map σ E₁ E₂)) := vector_prebundle.trivialization_of_mem_pretrivialization_atlas _ ⟨e₁, e₂, he₁, he₂, rfl⟩ instance _root_.bundle.continuous_linear_map.mem_trivialization_atlas : mem_trivialization_atlas (e₁.continuous_linear_map σ e₂ : - trivialization (F₁ →SL[σ] F₂) (π (bundle.continuous_linear_map σ F₁ E₁ F₂ E₂))) := + trivialization (F₁ →SL[σ] F₂) (π (F₁ →SL[σ] F₂) (bundle.continuous_linear_map σ E₁ E₂))) := { out := ⟨_, ⟨e₁, e₂, by apply_instance, by apply_instance, rfl⟩, rfl⟩ } variables {e₁ e₂} @@ -338,7 +312,7 @@ variables {e₁ e₂} rfl lemma trivialization.continuous_linear_map_apply - (p : total_space (bundle.continuous_linear_map σ F₁ E₁ F₂ E₂)) : + (p : total_space (F₁ →SL[σ] F₂) (bundle.continuous_linear_map σ E₁ E₂)) : e₁.continuous_linear_map σ e₂ p = ⟨p.1, (e₂.continuous_linear_map_at 𝕜₂ p.1 : _ →L[𝕜₂] _).comp (p.2.comp (e₁.symmL 𝕜₁ p.1 : F₁ →L[𝕜₁] E₁ p.1) : F₁ →SL[σ] E₂ p.1)⟩ := @@ -347,20 +321,20 @@ rfl omit he₁ he₂ lemma hom_trivialization_at_apply (x₀ : B) - (x : total_space (bundle.continuous_linear_map σ F₁ E₁ F₂ E₂)) : - trivialization_at (F₁ →SL[σ] F₂) (bundle.continuous_linear_map σ F₁ E₁ F₂ E₂) x₀ x = + (x : total_space (F₁ →SL[σ] F₂) (bundle.continuous_linear_map σ E₁ E₂)) : + trivialization_at (F₁ →SL[σ] F₂) (λ x, E₁ x →SL[σ] E₂ x) x₀ x = ⟨x.1, in_coordinates F₁ E₁ F₂ E₂ x₀ x.1 x₀ x.1 x.2⟩ := rfl @[simp, mfld_simps] lemma hom_trivialization_at_source (x₀ : B) : - (trivialization_at (F₁ →SL[σ] F₂) (bundle.continuous_linear_map σ F₁ E₁ F₂ E₂) x₀).source = - π (bundle.continuous_linear_map σ F₁ E₁ F₂ E₂) ⁻¹' + (trivialization_at (F₁ →SL[σ] F₂) (bundle.continuous_linear_map σ E₁ E₂) x₀).source = + π (F₁ →SL[σ] F₂) (bundle.continuous_linear_map σ E₁ E₂) ⁻¹' ((trivialization_at F₁ E₁ x₀).base_set ∩ (trivialization_at F₂ E₂ x₀).base_set) := rfl @[simp, mfld_simps] lemma hom_trivialization_at_target (x₀ : B) : - (trivialization_at (F₁ →SL[σ] F₂) (bundle.continuous_linear_map σ F₁ E₁ F₂ E₂) x₀).target = + (trivialization_at (F₁ →SL[σ] F₂) (λ x, E₁ x →SL[σ] E₂ x) x₀).target = ((trivialization_at F₁ E₁ x₀).base_set ∩ (trivialization_at F₂ E₂ x₀).base_set) ×ˢ set.univ := rfl From 728ef9dbb281241906f25cbeb30f90d83e0bb451 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Mon, 3 Jul 2023 06:18:05 +0000 Subject: [PATCH 1234/1291] chore(*): add mathlib4 synchronization comments (#19226) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.category.Mon.adjunctions` * `algebra.category.Semigroup.basic` * `algebraic_geometry.function_field` * `algebraic_geometry.morphisms.basic` * `algebraic_geometry.morphisms.open_immersion` * `algebraic_geometry.morphisms.universally_closed` * `analysis.complex.upper_half_plane.manifold` * `analysis.convex.cone.proper` * `category_theory.limits.constructions.over.basic` * `geometry.manifold.algebra.smooth_functions` * `geometry.manifold.derivation_bundle` * `number_theory.modular_forms.jacobi_theta.manifold` * `representation_theory.fdRep` * `representation_theory.invariants` * `ring_theory.witt_vector.isocrystal` * `set_theory.game.domineering` * `set_theory.game.short` * `set_theory.game.state` * `topology.homotopy.product` --- src/algebra/category/Mon/adjunctions.lean | 3 +++ src/algebra/category/Semigroup/basic.lean | 3 +++ src/algebraic_geometry/function_field.lean | 3 +++ src/algebraic_geometry/morphisms/basic.lean | 3 +++ src/algebraic_geometry/morphisms/open_immersion.lean | 3 +++ src/algebraic_geometry/morphisms/universally_closed.lean | 3 +++ src/analysis/complex/upper_half_plane/manifold.lean | 3 +++ src/analysis/convex/cone/proper.lean | 3 +++ src/category_theory/limits/constructions/over/basic.lean | 3 +++ src/geometry/manifold/algebra/smooth_functions.lean | 3 +++ src/geometry/manifold/derivation_bundle.lean | 3 +++ src/number_theory/modular_forms/jacobi_theta/manifold.lean | 3 +++ src/representation_theory/fdRep.lean | 3 +++ src/representation_theory/invariants.lean | 3 +++ src/ring_theory/witt_vector/isocrystal.lean | 3 +++ src/set_theory/game/domineering.lean | 3 +++ src/set_theory/game/short.lean | 3 +++ src/set_theory/game/state.lean | 3 +++ src/topology/homotopy/product.lean | 3 +++ 19 files changed, 57 insertions(+) diff --git a/src/algebra/category/Mon/adjunctions.lean b/src/algebra/category/Mon/adjunctions.lean index 54e756448f059..d3e500532d249 100644 --- a/src/algebra/category/Mon/adjunctions.lean +++ b/src/algebra/category/Mon/adjunctions.lean @@ -11,6 +11,9 @@ import algebra.free_monoid.basic /-! # Adjunctions regarding the category of monoids +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves the adjunction between adjoining a unit to a semigroup and the forgetful functor from monoids to semigroups. diff --git a/src/algebra/category/Semigroup/basic.lean b/src/algebra/category/Semigroup/basic.lean index 65bb2df05eb96..7e8666e59eda3 100644 --- a/src/algebra/category/Semigroup/basic.lean +++ b/src/algebra/category/Semigroup/basic.lean @@ -12,6 +12,9 @@ import category_theory.elementwise /-! # Category instances for has_mul, has_add, semigroup and add_semigroup +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We introduce the bundled categories: * `Magma` * `AddMagma` diff --git a/src/algebraic_geometry/function_field.lean b/src/algebraic_geometry/function_field.lean index 6e8f90ac361b0..628ccd684a434 100644 --- a/src/algebraic_geometry/function_field.lean +++ b/src/algebraic_geometry/function_field.lean @@ -8,6 +8,9 @@ import algebraic_geometry.properties /-! # Function field of integral schemes +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define the function field of an irreducible scheme as the stalk of the generic point. This is a field when the scheme is integral. diff --git a/src/algebraic_geometry/morphisms/basic.lean b/src/algebraic_geometry/morphisms/basic.lean index 322daeb6e9816..5ff036cda8b73 100644 --- a/src/algebraic_geometry/morphisms/basic.lean +++ b/src/algebraic_geometry/morphisms/basic.lean @@ -10,6 +10,9 @@ import category_theory.morphism_property /-! # Properties of morphisms between Schemes +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We provide the basic framework for talking about properties of morphisms between Schemes. A `morphism_property Scheme` is a predicate on morphisms between schemes, and an diff --git a/src/algebraic_geometry/morphisms/open_immersion.lean b/src/algebraic_geometry/morphisms/open_immersion.lean index 62465fea56bb8..4ae87dfbdce6e 100644 --- a/src/algebraic_geometry/morphisms/open_immersion.lean +++ b/src/algebraic_geometry/morphisms/open_immersion.lean @@ -10,6 +10,9 @@ import algebraic_geometry.morphisms.basic # Open immersions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A morphism is an open immersions if the underlying map of spaces is an open embedding `f : X ⟶ U ⊆ Y`, and the sheaf map `Y(V) ⟶ f _* X(V)` is an iso for each `V ⊆ U`. diff --git a/src/algebraic_geometry/morphisms/universally_closed.lean b/src/algebraic_geometry/morphisms/universally_closed.lean index 036ff1844d1d9..1d53bcffeb4ca 100644 --- a/src/algebraic_geometry/morphisms/universally_closed.lean +++ b/src/algebraic_geometry/morphisms/universally_closed.lean @@ -9,6 +9,9 @@ import topology.local_at_target /-! # Universally closed morphism +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A morphism of schemes `f : X ⟶ Y` is universally closed if `X ×[Y] Y' ⟶ Y'` is a closed map for all base change `Y' ⟶ Y`. diff --git a/src/analysis/complex/upper_half_plane/manifold.lean b/src/analysis/complex/upper_half_plane/manifold.lean index b5a8b7ec6de4b..0c070195cc5dd 100644 --- a/src/analysis/complex/upper_half_plane/manifold.lean +++ b/src/analysis/complex/upper_half_plane/manifold.lean @@ -8,6 +8,9 @@ import geometry.manifold.cont_mdiff_mfderiv /-! # Manifold structure on the upper half plane. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define the complex manifold structure on the upper half-plane. -/ diff --git a/src/analysis/convex/cone/proper.lean b/src/analysis/convex/cone/proper.lean index 9babbda4c6ef3..8a76a074f3ae9 100644 --- a/src/analysis/convex/cone/proper.lean +++ b/src/analysis/convex/cone/proper.lean @@ -9,6 +9,9 @@ import analysis.inner_product_space.adjoint /-! # Proper cones +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define a proper cone as a nonempty, closed, convex cone. Proper cones are used in defining conic programs which generalize linear programs. A linear program is a conic program for the positive cone. We then prove Farkas' lemma for conic programs following the proof in the reference below. diff --git a/src/category_theory/limits/constructions/over/basic.lean b/src/category_theory/limits/constructions/over/basic.lean index e4085355f2824..9e71f7f3ed156 100644 --- a/src/category_theory/limits/constructions/over/basic.lean +++ b/src/category_theory/limits/constructions/over/basic.lean @@ -12,6 +12,9 @@ import category_theory.limits.constructions.equalizers /-! # Limits in the over category +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Declare instances for limits in the over category: If `C` has finite wide pullbacks, `over B` has finite limits, and if `C` has arbitrary wide pullbacks then `over B` has limits. -/ diff --git a/src/geometry/manifold/algebra/smooth_functions.lean b/src/geometry/manifold/algebra/smooth_functions.lean index 932b3e776ebf1..47e08e8c88e99 100644 --- a/src/geometry/manifold/algebra/smooth_functions.lean +++ b/src/geometry/manifold/algebra/smooth_functions.lean @@ -9,6 +9,9 @@ import geometry.manifold.algebra.structures /-! # Algebraic structures over smooth functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we define instances of algebraic structures over smooth functions. -/ diff --git a/src/geometry/manifold/derivation_bundle.lean b/src/geometry/manifold/derivation_bundle.lean index 71f103747b268..d4ad6affb65b6 100644 --- a/src/geometry/manifold/derivation_bundle.lean +++ b/src/geometry/manifold/derivation_bundle.lean @@ -11,6 +11,9 @@ import ring_theory.derivation.basic # Derivation bundle +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define the derivations at a point of a manifold on the algebra of smooth fuctions. Moreover, we define the differential of a function in terms of derivations. diff --git a/src/number_theory/modular_forms/jacobi_theta/manifold.lean b/src/number_theory/modular_forms/jacobi_theta/manifold.lean index 67094c9757bca..81a9584cd7556 100644 --- a/src/number_theory/modular_forms/jacobi_theta/manifold.lean +++ b/src/number_theory/modular_forms/jacobi_theta/manifold.lean @@ -9,6 +9,9 @@ import analysis.complex.upper_half_plane.manifold /-! # Manifold differentiability of the Jacobi's theta function +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we reformulate differentiability of the Jacobi's theta function in terms of manifold differentiability. diff --git a/src/representation_theory/fdRep.lean b/src/representation_theory/fdRep.lean index 213f589b738dd..70cef6ac6b148 100644 --- a/src/representation_theory/fdRep.lean +++ b/src/representation_theory/fdRep.lean @@ -11,6 +11,9 @@ import representation_theory.basic /-! # `fdRep k G` is the category of finite dimensional `k`-linear representations of `G`. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + If `V : fdRep k G`, there is a coercion that allows you to treat `V` as a type, and this type comes equipped with `module k V` and `finite_dimensional k V` instances. Also `V.ρ` gives the homomorphism `G →* (V →ₗ[k] V)`. diff --git a/src/representation_theory/invariants.lean b/src/representation_theory/invariants.lean index 465fc5df768ae..5a9380d2cc6c3 100644 --- a/src/representation_theory/invariants.lean +++ b/src/representation_theory/invariants.lean @@ -9,6 +9,9 @@ import representation_theory.fdRep /-! # Subspace of invariants a group representation +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file introduces the subspace of invariants of a group representation and proves basic results about it. The main tool used is the average of all elements of the group, seen as an element of diff --git a/src/ring_theory/witt_vector/isocrystal.lean b/src/ring_theory/witt_vector/isocrystal.lean index bc06cef87421a..02374a808e1f4 100644 --- a/src/ring_theory/witt_vector/isocrystal.lean +++ b/src/ring_theory/witt_vector/isocrystal.lean @@ -10,6 +10,9 @@ import ring_theory.witt_vector.frobenius_fraction_field ## F-isocrystals over a perfect field +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + When `k` is an integral domain, so is `𝕎 k`, and we can consider its field of fractions `K(p, k)`. The endomorphism `witt_vector.frobenius` lifts to `φ : K(p, k) → K(p, k)`; if `k` is perfect, `φ` is an automorphism. diff --git a/src/set_theory/game/domineering.lean b/src/set_theory/game/domineering.lean index 3190d0fa99854..43958be820b70 100644 --- a/src/set_theory/game/domineering.lean +++ b/src/set_theory/game/domineering.lean @@ -8,6 +8,9 @@ import set_theory.game.state /-! # Domineering as a combinatorial game. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define the game of Domineering, played on a chessboard of arbitrary shape (possibly even disconnected). Left moves by placing a domino vertically, while Right moves by placing a domino horizontally. diff --git a/src/set_theory/game/short.lean b/src/set_theory/game/short.lean index 8d2bb319530e5..6b5867a25770c 100644 --- a/src/set_theory/game/short.lean +++ b/src/set_theory/game/short.lean @@ -10,6 +10,9 @@ import set_theory.game.birthday /-! # Short games +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A combinatorial game is `short` [Conway, ch.9][conway2001] if it has only finitely many positions. In particular, this means there is a finite set of moves at every point. diff --git a/src/set_theory/game/state.lean b/src/set_theory/game/state.lean index 14f118bc5685d..27f9ca055dce5 100644 --- a/src/set_theory/game/state.lean +++ b/src/set_theory/game/state.lean @@ -8,6 +8,9 @@ import set_theory.game.short /-! # Games described via "the state of the board". +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We provide a simple mechanism for constructing combinatorial (pre-)games, by describing "the state of the board", and providing an upper bound on the number of turns remaining. diff --git a/src/topology/homotopy/product.lean b/src/topology/homotopy/product.lean index 06b9d0449ed71..bc5ccd0df73f3 100644 --- a/src/topology/homotopy/product.lean +++ b/src/topology/homotopy/product.lean @@ -9,6 +9,9 @@ import topology.homotopy.path /-! # Product of homotopies +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we introduce definitions for the product of homotopies. We show that the products of relative homotopies are still relative homotopies. Finally, we specialize to the case From 5dc6092d09e5e489106865241986f7f2ad28d4c8 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Tue, 4 Jul 2023 07:03:10 +0000 Subject: [PATCH 1235/1291] chore(topology/sheaves): revert universe generalizations from #19153 (#19230) This reverts commit 13361559d66b84f80b6d5a1c4a26aa5054766725. These are just too difficult to forward port as is because of the `max u v =?= max u ?v` issue https://github.com/leanprover/lean4/issues/2297. We have another candidate approach to this, using a new `UnivLE` typeclass, and I would prefer if we investigated that without the pressure of the port at the same time. This will delay @hrmacbeth's plans to define meromorphic functions, perhaps. Co-authored-by: Scott Morrison --- .../morphisms/quasi_compact.lean | 4 +- .../morphisms/quasi_separated.lean | 6 +-- src/algebraic_geometry/ringed_space.lean | 4 +- src/algebraic_geometry/structure_sheaf.lean | 4 +- src/category_theory/limits/shapes/types.lean | 19 +-------- src/topology/sheaves/forget.lean | 28 ++++++------- src/topology/sheaves/local_predicate.lean | 40 ++++++++----------- src/topology/sheaves/presheaf.lean | 20 +++++----- .../sheaves/presheaf_of_functions.lean | 8 ++-- .../sheaf_condition/unique_gluing.lean | 29 +++++++------- src/topology/sheaves/sheaf_of_functions.lean | 8 ++-- src/topology/sheaves/stalks.lean | 6 +-- 12 files changed, 74 insertions(+), 102 deletions(-) diff --git a/src/algebraic_geometry/morphisms/quasi_compact.lean b/src/algebraic_geometry/morphisms/quasi_compact.lean index 5da4cb3432cfd..b3281898a3059 100644 --- a/src/algebraic_geometry/morphisms/quasi_compact.lean +++ b/src/algebraic_geometry/morphisms/quasi_compact.lean @@ -298,7 +298,7 @@ end /-- If `x : Γ(X, U)` is zero on `D(f)` for some `f : Γ(X, U)`, and `U` is quasi-compact, then `f ^ n * x = 0` for some `n`. -/ -lemma exists_pow_mul_eq_zero_of_res_basic_open_eq_zero_of_is_compact (X : Scheme.{u}) +lemma exists_pow_mul_eq_zero_of_res_basic_open_eq_zero_of_is_compact (X : Scheme) {U : opens X.carrier} (hU : is_compact U.1) (x f : X.presheaf.obj (op U)) (H : x |_ X.basic_open f = 0) : ∃ n : ℕ, f ^ n * x = 0 := @@ -322,7 +322,7 @@ begin use finset.univ.sup n, suffices : ∀ (i : s), X.presheaf.map (hom_of_le (h₁ i)).op (f ^ (finset.univ.sup n) * x) = 0, { subst e, - apply Top.sheaf.eq_of_locally_eq.{(u+1) u} X.sheaf (λ (i : s), (i : opens X.carrier)), + apply X.sheaf.eq_of_locally_eq (λ (i : s), (i : opens X.carrier)), intro i, rw map_zero, apply this }, diff --git a/src/algebraic_geometry/morphisms/quasi_separated.lean b/src/algebraic_geometry/morphisms/quasi_separated.lean index f00888c66f0ba..8b75d385dcf82 100644 --- a/src/algebraic_geometry/morphisms/quasi_separated.lean +++ b/src/algebraic_geometry/morphisms/quasi_separated.lean @@ -345,7 +345,7 @@ begin exact e end -lemma exists_eq_pow_mul_of_is_compact_of_is_quasi_separated (X : Scheme.{u}) +lemma exists_eq_pow_mul_of_is_compact_of_is_quasi_separated (X : Scheme) (U : opens X.carrier) (hU : is_compact U.1) (hU' : is_quasi_separated U.1) (f : X.presheaf.obj (op U)) (x : X.presheaf.obj (op $ X.basic_open f)) : ∃ (n : ℕ) (y : X.presheaf.obj (op U)), y |_ X.basic_open f = (f |_ X.basic_open f) ^ n * x := @@ -401,7 +401,7 @@ begin ((X.presheaf.map (hom_of_le le_sup_left).op f) ^ (finset.univ.sup n + n₂) * y₁) = X.presheaf.map (hom_of_le $ inf_le_right).op ((X.presheaf.map (hom_of_le le_sup_right).op f) ^ (finset.univ.sup n + n₁) * y₂), - { fapply Top.sheaf.eq_of_locally_eq'.{(u+1) u} X.sheaf (λ i : s, i.1.1), + { fapply X.sheaf.eq_of_locally_eq' (λ i : s, i.1.1), { refine λ i, hom_of_le _, erw hs, exact le_supr _ _ }, { exact le_of_eq hs }, { intro i, @@ -419,7 +419,7 @@ begin -- By the sheaf condition, since `f ^ (n + n₂) * y₁ = f ^ (n + n₁) * y₂`, it can be glued into -- the desired section on `S ∪ U`. use (X.sheaf.obj_sup_iso_prod_eq_locus S U.1).inv ⟨⟨_ * _, _ * _⟩, this⟩, - refine Top.sheaf.eq_of_locally_eq₂.{(u+1) u} X.sheaf + refine X.sheaf.eq_of_locally_eq₂ (hom_of_le (_ : X.basic_open (X.presheaf.map (hom_of_le le_sup_left).op f) ≤ _)) (hom_of_le (_ : X.basic_open (X.presheaf.map (hom_of_le le_sup_right).op f) ≤ _)) _ _ _ _ _, { rw X.basic_open_res, exact inf_le_right }, diff --git a/src/algebraic_geometry/ringed_space.lean b/src/algebraic_geometry/ringed_space.lean index 268e8c7f83277..ff9354559ca56 100644 --- a/src/algebraic_geometry/ringed_space.lean +++ b/src/algebraic_geometry/ringed_space.lean @@ -76,7 +76,7 @@ begin -- Let `g x` denote the inverse of `f` in `U x`. choose g hg using λ x : U, is_unit.exists_right_inv (h_unit x), -- We claim that these local inverses glue together to a global inverse of `f`. - obtain ⟨gl, gl_spec, -⟩ := Top.sheaf.exists_unique_gluing'.{(v+1) v} X.sheaf V U iVU hcover g _, + obtain ⟨gl, gl_spec, -⟩ := X.sheaf.exists_unique_gluing' V U iVU hcover g _, swap, { intros x y, apply section_ext X.sheaf (V x ⊓ V y), @@ -89,7 +89,7 @@ begin congr_arg (X.presheaf.germ (⟨z, hzVy⟩ : V y)) (hg y), ring_hom.map_one, ring_hom.map_one] }, apply is_unit_of_mul_eq_one f gl, - apply Top.sheaf.eq_of_locally_eq'.{(v+1) v} X.sheaf V U iVU hcover, + apply X.sheaf.eq_of_locally_eq' V U iVU hcover, intro i, rw [ring_hom.map_one, ring_hom.map_mul, gl_spec], exact hg i, diff --git a/src/algebraic_geometry/structure_sheaf.lean b/src/algebraic_geometry/structure_sheaf.lean index 799927233cb81..1f2848001ed88 100644 --- a/src/algebraic_geometry/structure_sheaf.lean +++ b/src/algebraic_geometry/structure_sheaf.lean @@ -772,10 +772,10 @@ begin rw to_basic_open_mk', -- Since the structure sheaf is a sheaf, we can show the desired equality locally. - -- Annoyingly, `sheaf.eq_of_locally_eq'` requires an open cover indexed by a *type*, so we need to + -- Annoyingly, `sheaf.eq_of_locally_eq` requires an open cover indexed by a *type*, so we need to -- coerce our finset `t` to a type first. let tt := ((t : set (basic_open f)) : Type u), - apply Top.sheaf.eq_of_locally_eq'.{(u+1) u} (structure_sheaf R) + apply (structure_sheaf R).eq_of_locally_eq' (λ i : tt, basic_open (h i)) (basic_open f) (λ i : tt, iDh i), { -- This feels a little redundant, since already have `ht_cover` as a hypothesis -- Unfortunately, `ht_cover` uses a bounded union over the set `t`, while here we have the diff --git a/src/category_theory/limits/shapes/types.lean b/src/category_theory/limits/shapes/types.lean index 6439eb4210fb4..4666f03ce1a24 100644 --- a/src/category_theory/limits/shapes/types.lean +++ b/src/category_theory/limits/shapes/types.lean @@ -51,28 +51,13 @@ local attribute [tidy] tactic.discrete_cases /-- A restatement of `types.lift_π_apply` that uses `pi.π` and `pi.lift`. -/ @[simp] lemma pi_lift_π_apply - {β : Type v} (f : β → Type max v u) {P : Type max v u} (s : Π b, P ⟶ f b) (b : β) (x : P) : - (pi.π f b : (∏ f) → f b) (@pi.lift β _ _ f _ P s x) = s b x := -congr_fun (limit.lift_π (fan.mk P s) ⟨b⟩) x - -/-- A restatement of `types.lift_π_apply` that uses `pi.π` and `pi.lift`, -with specialized universes. -/ -@[simp] -lemma pi_lift_π_apply' - {β : Type v} (f : β → Type v) {P : Type v} (s : Π b, P ⟶ f b) (b : β) (x : P) : + {β : Type u} (f : β → Type u) {P : Type u} (s : Π b, P ⟶ f b) (b : β) (x : P) : (pi.π f b : (∏ f) → f b) (@pi.lift β _ _ f _ P s x) = s b x := congr_fun (limit.lift_π (fan.mk P s) ⟨b⟩) x /-- A restatement of `types.map_π_apply` that uses `pi.π` and `pi.map`. -/ @[simp] -lemma pi_map_π_apply {β : Type v} {f g : β → Type max v u} (α : Π j, f j ⟶ g j) (b : β) (x) : - (pi.π g b : (∏ g) → g b) (pi.map α x) = α b ((pi.π f b : (∏ f) → f b) x) := -limit.map_π_apply _ _ _ - -/-- A restatement of `types.map_π_apply` that uses `pi.π` and `pi.map`, -with specialized universes. -/ -@[simp] -lemma pi_map_π_apply' {β : Type v} {f g : β → Type v} (α : Π j, f j ⟶ g j) (b : β) (x) : +lemma pi_map_π_apply {β : Type u} {f g : β → Type u} (α : Π j, f j ⟶ g j) (b : β) (x) : (pi.π g b : (∏ g) → g b) (pi.map α x) = α b ((pi.π f b : (∏ f) → f b) x) := limit.map_π_apply _ _ _ diff --git a/src/topology/sheaves/forget.lean b/src/topology/sheaves/forget.lean index f6fbde37a2ce5..3fb52ea35b299 100644 --- a/src/topology/sheaves/forget.lean +++ b/src/topology/sheaves/forget.lean @@ -42,13 +42,13 @@ namespace sheaf_condition open sheaf_condition_equalizer_products -universes v u₁ u₂ w +universes v u₁ u₂ -variables {C : Type u₁} [category.{v} C] [has_limits_of_size.{w w} C] -variables {D : Type u₂} [category.{v} D] [has_limits_of_size.{w w} D] -variables (G : C ⥤ D) [preserves_limits_of_size.{w w} G] -variables {X : Top.{w}} (F : presheaf C X) -variables {ι : Type w} (U : ι → opens X) +variables {C : Type u₁} [category.{v} C] [has_limits C] +variables {D : Type u₂} [category.{v} D] [has_limits D] +variables (G : C ⥤ D) [preserves_limits G] +variables {X : Top.{v}} (F : presheaf C X) +variables {ι : Type v} (U : ι → opens X) local attribute [reducible] diagram left_res right_res @@ -57,7 +57,7 @@ When `G` preserves limits, the sheaf condition diagram for `F` composed with `G` naturally isomorphic to the sheaf condition diagram for `F ⋙ G`. -/ def diagram_comp_preserves_limits : - diagram F U ⋙ G ≅ diagram.{w} (F ⋙ G) U := + diagram F U ⋙ G ≅ diagram.{v} (F ⋙ G) U := begin fapply nat_iso.of_components, rintro ⟨j⟩, @@ -85,7 +85,7 @@ When `G` preserves limits, the image under `G` of the sheaf condition fork for ` is the sheaf condition fork for `F ⋙ G`, postcomposed with the inverse of the natural isomorphism `diagram_comp_preserves_limits`. -/ -def map_cone_fork : G.map_cone (fork.{w} F U) ≅ +def map_cone_fork : G.map_cone (fork.{v} F U) ≅ (cones.postcompose (diagram_comp_preserves_limits G F U).inv).obj (fork (F ⋙ G) U) := cones.ext (iso.refl _) (λ j, begin @@ -102,17 +102,16 @@ end) end sheaf_condition -universes v u₁ u₂ w +universes v u₁ u₂ open sheaf_condition sheaf_condition_equalizer_products variables {C : Type u₁} [category.{v} C] {D : Type u₂} [category.{v} D] variables (G : C ⥤ D) variables [reflects_isomorphisms G] -variables [has_limits_of_size.{w w} C] [has_limits_of_size.{w w} D] - [preserves_limits_of_size.{w w} G] +variables [has_limits C] [has_limits D] [preserves_limits G] -variables {X : Top.{w}} (F : presheaf C X) +variables {X : Top.{v}} (F : presheaf C X) /-- If `G : C ⥤ D` is a functor which reflects isomorphisms and preserves limits @@ -172,7 +171,7 @@ begin -- image under `G` of the equalizer cone for the sheaf condition diagram. let c := fork (F ⋙ G) U, obtain ⟨hc⟩ := S U, - let d := G.map_cone (equalizer.fork (left_res.{w} F U) (right_res F U)), + let d := G.map_cone (equalizer.fork (left_res.{v} F U) (right_res F U)), letI := preserves_smallest_limits_of_preserves_limits G, have hd : is_limit d := preserves_limit.preserves (limit.is_limit _), -- Since both of these are limit cones @@ -183,8 +182,7 @@ begin -- introduced above. let d' := (cones.postcompose (diagram_comp_preserves_limits G F U).hom).obj d, have hd' : is_limit d' := - (is_limit.postcompose_hom_equiv - (diagram_comp_preserves_limits.{v u₁ u₂ w} G F U : _) d).symm hd, + (is_limit.postcompose_hom_equiv (diagram_comp_preserves_limits G F U : _) d).symm hd, -- Now everything works: we verify that `f` really is a morphism between these cones: let f' : c ⟶ d' := fork.mk_hom (G.map f) diff --git a/src/topology/sheaves/local_predicate.lean b/src/topology/sheaves/local_predicate.lean index 3b682723589bf..04659ff17fd1b 100644 --- a/src/topology/sheaves/local_predicate.lean +++ b/src/topology/sheaves/local_predicate.lean @@ -38,12 +38,12 @@ to the types in the ambient type family. We give conditions sufficient to show that this map is injective and/or surjective. -/ -universes v w +universe v noncomputable theory variables {X : Top.{v}} -variables (T : X → Type w) +variables (T : X → Type v) open topological_space open opposite @@ -69,12 +69,12 @@ variables (X) Continuity is a "prelocal" predicate on functions to a fixed topological space `T`. -/ @[simps] -def continuous_prelocal (T : Top.{w}) : prelocal_predicate (λ x : X, T) := +def continuous_prelocal (T : Top.{v}) : prelocal_predicate (λ x : X, T) := { pred := λ U f, continuous f, res := λ U V i f h, continuous.comp h (opens.open_embedding_of_le i.le).continuous, } /-- Satisfying the inhabited linter. -/ -instance inhabited_prelocal_predicate (T : Top.{w}) : inhabited (prelocal_predicate (λ x : X, T)) := +instance inhabited_prelocal_predicate (T : Top.{v}) : inhabited (prelocal_predicate (λ x : X, T)) := ⟨continuous_prelocal X T⟩ variables {X} @@ -150,7 +150,7 @@ lemma prelocal_predicate.sheafify_of {T : X → Type v} {P : prelocal_predicate The subpresheaf of dependent functions on `X` satisfying the "pre-local" predicate `P`. -/ @[simps] -def subpresheaf_to_Types (P : prelocal_predicate T) : presheaf (Type (max v w)) X := +def subpresheaf_to_Types (P : prelocal_predicate T) : presheaf (Type v) X := { obj := λ U, { f : Π x : unop U, T x // P.pred f }, map := λ U V i f, ⟨λ x, f.1 (i.unop x), P.res i.unop f.1 f.2⟩ }. @@ -211,27 +211,22 @@ end subpresheaf_to_Types The subsheaf of the sheaf of all dependently typed functions satisfying the local predicate `P`. -/ @[simps] -def subsheaf_to_Types (P : local_predicate T) : sheaf (Type max v w) X := +def subsheaf_to_Types (P : local_predicate T) : sheaf (Type v) X := ⟨subpresheaf_to_Types P.to_prelocal_predicate, subpresheaf_to_Types.is_sheaf P⟩ --- TODO There is further universe generalization to do here, --- but it depends on more difficult universe generalization in `topology.sheaves.stalks`. --- The aim is to replace `T'` with the more general `T` below. -variables {T' : X → Type v} - /-- There is a canonical map from the stalk to the original fiber, given by evaluating sections. -/ -def stalk_to_fiber (P : local_predicate T') (x : X) : - (subsheaf_to_Types P).presheaf.stalk x ⟶ T' x := +def stalk_to_fiber (P : local_predicate T) (x : X) : + (subsheaf_to_Types P).presheaf.stalk x ⟶ T x := begin refine colimit.desc _ - { X := T' x, ι := { app := λ U f, _, naturality' := _ } }, + { X := T x, ι := { app := λ U f, _, naturality' := _ } }, { exact f.1 ⟨x, (unop U).2⟩, }, { tidy, } end -@[simp] lemma stalk_to_fiber_germ (P : local_predicate T') (U : opens X) (x : U) (f) : +@[simp] lemma stalk_to_fiber_germ (P : local_predicate T) (U : opens X) (x : U) (f) : stalk_to_fiber P x ((subsheaf_to_Types P).presheaf.germ x f) = f.1 x := begin dsimp [presheaf.germ, stalk_to_fiber], @@ -244,8 +239,8 @@ end The `stalk_to_fiber` map is surjective at `x` if every point in the fiber `T x` has an allowed section passing through it. -/ -lemma stalk_to_fiber_surjective (P : local_predicate T') (x : X) - (w : ∀ (t : T' x), ∃ (U : open_nhds x) (f : Π y : U.1, T' y) (h : P.pred f), f ⟨x, U.2⟩ = t) : +lemma stalk_to_fiber_surjective (P : local_predicate T) (x : X) + (w : ∀ (t : T x), ∃ (U : open_nhds x) (f : Π y : U.1, T y) (h : P.pred f), f ⟨x, U.2⟩ = t) : function.surjective (stalk_to_fiber P x) := λ t, begin @@ -259,16 +254,16 @@ end The `stalk_to_fiber` map is injective at `x` if any two allowed sections which agree at `x` agree on some neighborhood of `x`. -/ -lemma stalk_to_fiber_injective (P : local_predicate T') (x : X) - (w : ∀ (U V : open_nhds x) (fU : Π y : U.1, T' y) (hU : P.pred fU) - (fV : Π y : V.1, T' y) (hV : P.pred fV) (e : fU ⟨x, U.2⟩ = fV ⟨x, V.2⟩), +lemma stalk_to_fiber_injective (P : local_predicate T) (x : X) + (w : ∀ (U V : open_nhds x) (fU : Π y : U.1, T y) (hU : P.pred fU) + (fV : Π y : V.1, T y) (hV : P.pred fV) (e : fU ⟨x, U.2⟩ = fV ⟨x, V.2⟩), ∃ (W : open_nhds x) (iU : W ⟶ U) (iV : W ⟶ V), ∀ (w : W.1), fU (iU w : U.1) = fV (iV w : V.1)) : function.injective (stalk_to_fiber P x) := λ tU tV h, begin -- We promise to provide all the ingredients of the proof later: let Q : - ∃ (W : (open_nhds x)ᵒᵖ) (s : Π w : (unop W).1, T' w) (hW : P.pred s), + ∃ (W : (open_nhds x)ᵒᵖ) (s : Π w : (unop W).1, T w) (hW : P.pred s), tU = (subsheaf_to_Types P).presheaf.germ ⟨x, (unop W).2⟩ ⟨s, hW⟩ ∧ tV = (subsheaf_to_Types P).presheaf.germ ⟨x, (unop W).2⟩ ⟨s, hW⟩ := _, { choose W s hW e using Q, @@ -288,9 +283,6 @@ begin colimit_sound iV.op (subtype.eq (funext w).symm)⟩, }, end --- TODO to continue generalizing universes here, we will need to deal with the issue noted --- at `presheaf_to_Top` (universe generalization for the Yoneda embedding). - /-- Some repackaging: the presheaf of functions satisfying `continuous_prelocal` is just the same thing as diff --git a/src/topology/sheaves/presheaf.lean b/src/topology/sheaves/presheaf.lean index ee4af1ab86ca9..e4c312be03265 100644 --- a/src/topology/sheaves/presheaf.lean +++ b/src/topology/sheaves/presheaf.lean @@ -221,7 +221,7 @@ def pushforward_map {X Y : Top.{w}} (f : X ⟶ Y) {ℱ 𝒢 : X.presheaf C} (α open category_theory.limits section pullback -variable [has_colimits_of_size.{w w} C] +variable [has_colimits C] noncomputable theory /-- @@ -231,17 +231,17 @@ This is defined in terms of left Kan extensions, which is just a fancy way of sa "take the colimits over the open sets whose preimage contains U". -/ @[simps] -def pullback_obj {X Y : Top.{w}} (f : X ⟶ Y) (ℱ : Y.presheaf C) : X.presheaf C := +def pullback_obj {X Y : Top.{v}} (f : X ⟶ Y) (ℱ : Y.presheaf C) : X.presheaf C := (Lan (opens.map f).op).obj ℱ /-- Pulling back along continuous maps is functorial. -/ -def pullback_map {X Y : Top.{w}} (f : X ⟶ Y) {ℱ 𝒢 : Y.presheaf C} (α : ℱ ⟶ 𝒢) : +def pullback_map {X Y : Top.{v}} (f : X ⟶ Y) {ℱ 𝒢 : Y.presheaf C} (α : ℱ ⟶ 𝒢) : pullback_obj f ℱ ⟶ pullback_obj f 𝒢 := (Lan (opens.map f).op).map α /-- If `f '' U` is open, then `f⁻¹ℱ U ≅ ℱ (f '' U)`. -/ @[simps] -def pullback_obj_obj_of_image_open {X Y : Top.{w}} (f : X ⟶ Y) (ℱ : Y.presheaf C) (U : opens X) +def pullback_obj_obj_of_image_open {X Y : Top.{v}} (f : X ⟶ Y) (ℱ : Y.presheaf C) (U : opens X) (H : is_open (f '' U)) : (pullback_obj f ℱ).obj (op U) ≅ ℱ.obj (op ⟨_, H⟩) := begin let x : costructured_arrow (opens.map f).op (op U) := begin @@ -263,7 +263,7 @@ begin end namespace pullback -variables {X Y : Top.{w}} (ℱ : Y.presheaf C) +variables {X Y : Top.{v}} (ℱ : Y.presheaf C) /-- The pullback along the identity is isomorphic to the original presheaf. -/ def id : pullback_obj (𝟙 _) ℱ ≅ ℱ := @@ -366,30 +366,30 @@ by simpa [pushforward_to_of_iso, equivalence.to_adjunction] end iso -variables (C) [has_colimits_of_size.{w w} C] +variables (C) [has_colimits C] /-- Pullback a presheaf on `Y` along a continuous map `f : X ⟶ Y`, obtaining a presheaf on `X`. -/ @[simps map_app] -def pullback {X Y : Top.{w}} (f : X ⟶ Y) : Y.presheaf C ⥤ X.presheaf C := Lan (opens.map f).op +def pullback {X Y : Top.{v}} (f : X ⟶ Y) : Y.presheaf C ⥤ X.presheaf C := Lan (opens.map f).op @[simp] lemma pullback_obj_eq_pullback_obj {C} [category C] [has_colimits C] {X Y : Top.{w}} (f : X ⟶ Y) (ℱ : Y.presheaf C) : (pullback C f).obj ℱ = pullback_obj f ℱ := rfl /-- The pullback and pushforward along a continuous map are adjoint to each other. -/ @[simps unit_app_app counit_app_app] -def pushforward_pullback_adjunction {X Y : Top.{w}} (f : X ⟶ Y) : +def pushforward_pullback_adjunction {X Y : Top.{v}} (f : X ⟶ Y) : pullback C f ⊣ pushforward C f := Lan.adjunction _ _ /-- Pulling back along a homeomorphism is the same as pushing forward along its inverse. -/ -def pullback_hom_iso_pushforward_inv {X Y : Top.{w}} (H : X ≅ Y) : +def pullback_hom_iso_pushforward_inv {X Y : Top.{v}} (H : X ≅ Y) : pullback C H.hom ≅ pushforward C H.inv := adjunction.left_adjoint_uniq (pushforward_pullback_adjunction C H.hom) (presheaf_equiv_of_iso C H.symm).to_adjunction /-- Pulling back along the inverse of a homeomorphism is the same as pushing forward along it. -/ -def pullback_inv_iso_pushforward_hom {X Y : Top.{w}} (H : X ≅ Y) : +def pullback_inv_iso_pushforward_hom {X Y : Top.{v}} (H : X ≅ Y) : pullback C H.inv ≅ pushforward C H.hom := adjunction.left_adjoint_uniq (pushforward_pullback_adjunction C H.inv) diff --git a/src/topology/sheaves/presheaf_of_functions.lean b/src/topology/sheaves/presheaf_of_functions.lean index a6c432a1dea51..07212f6643439 100644 --- a/src/topology/sheaves/presheaf_of_functions.lean +++ b/src/topology/sheaves/presheaf_of_functions.lean @@ -28,7 +28,7 @@ We construct some simple examples of presheaves of functions on a topological sp is the presheaf of rings of continuous complex-valued functions on `X`. -/ -universes v u w +universes v u open category_theory open topological_space @@ -42,7 +42,7 @@ variables (X : Top.{v}) The presheaf of dependently typed functions on `X`, with fibres given by a type family `T`. There is no requirement that the functions are continuous, here. -/ -def presheaf_to_Types (T : X → Type w) : X.presheaf (Type (max v w)) := +def presheaf_to_Types (T : X → Type v) : X.presheaf (Type v) := { obj := λ U, Π x : (unop U), T x, map := λ U V i g, λ (x : unop V), g (i.unop x), map_id' := λ U, by { ext g ⟨x, hx⟩, refl }, @@ -68,7 +68,7 @@ There is no requirement that the functions are continuous, here. -- We don't use `@[simps]` to generate the projection lemmas here, -- as it turns out to be useful to have `presheaf_to_Type_map` -- written as an equality of functions (rather than being applied to some argument). -def presheaf_to_Type (T : Type w) : X.presheaf (Type max v w) := +def presheaf_to_Type (T : Type v) : X.presheaf (Type v) := { obj := λ U, (unop U) → T, map := λ U V i g, g ∘ i.unop, map_id' := λ U, by { ext g ⟨x, hx⟩, refl }, @@ -86,8 +86,6 @@ rfl /-- The presheaf of continuous functions on `X` with values in fixed target topological space `T`. -/ --- TODO it may prove useful to generalize the universes here, --- but the definition would need to change. def presheaf_to_Top (T : Top.{v}) : X.presheaf (Type v) := (opens.to_Top X).op ⋙ (yoneda.obj T) diff --git a/src/topology/sheaves/sheaf_condition/unique_gluing.lean b/src/topology/sheaves/sheaf_condition/unique_gluing.lean index 543920fcc9b69..2cbe750fc2433 100644 --- a/src/topology/sheaves/sheaf_condition/unique_gluing.lean +++ b/src/topology/sheaves/sheaf_condition/unique_gluing.lean @@ -48,9 +48,9 @@ open topological_space open topological_space.opens open opposite -universes u v w +universes u v -variables {C : Type u} [category.{max w v} C] [concrete_category.{max w v} C] +variables {C : Type u} [category.{v} C] [concrete_category.{v} C] namespace Top @@ -60,7 +60,7 @@ section local attribute [instance] concrete_category.has_coe_to_sort concrete_category.has_coe_to_fun -variables {X : Top.{w}} (F : presheaf C X) {ι : Type w} (U : ι → opens X) +variables {X : Top.{v}} (F : presheaf C X) {ι : Type v} (U : ι → opens X) /-- A family of sections `sf` is compatible, if the restrictions of `sf i` and `sf j` to `U i ⊓ U j` @@ -85,14 +85,14 @@ We prove this to be equivalent to the usual one below in `is_sheaf_iff_is_sheaf_unique_gluing` -/ def is_sheaf_unique_gluing : Prop := -∀ ⦃ι : Type w⦄ (U : ι → opens X) (sf : Π i : ι, F.obj (op (U i))), +∀ ⦃ι : Type v⦄ (U : ι → opens X) (sf : Π i : ι, F.obj (op (U i))), is_compatible F U sf → ∃! s : F.obj (op (supr U)), is_gluing F U sf s end section type_valued -variables {X : Top.{w}} (F : presheaf (Type max w v) X) {ι : Type w} (U : ι → opens X) +variables {X : Top.{v}} (F : presheaf (Type v) X) {ι : Type v} (U : ι → opens X) /-- For presheaves of types, terms of `pi_opens F U` are just families of sections. @@ -100,7 +100,7 @@ For presheaves of types, terms of `pi_opens F U` are just families of sections. def pi_opens_iso_sections_family : pi_opens F U ≅ Π i : ι, F.obj (op (U i)) := limits.is_limit.cone_point_unique_up_to_iso (limit.is_limit (discrete.functor (λ i : ι, F.obj (op (U i))))) - ((types.product_limit_cone.{w (max w v)} (λ i : ι, F.obj (op (U i)))).is_limit) + ((types.product_limit_cone.{v v} (λ i : ι, F.obj (op (U i)))).is_limit) /-- Under the isomorphism `pi_opens_iso_sections_family`, compatibility of sections is the same @@ -112,8 +112,8 @@ lemma compatible_iff_left_res_eq_right_res (sf : pi_opens F U) : begin split ; intros h, { ext ⟨i, j⟩, - rw [left_res, types.limit.lift_π_apply, fan.mk_π_app, - right_res, types.limit.lift_π_apply, fan.mk_π_app], + rw [left_res, types.limit.lift_π_apply', fan.mk_π_app, + right_res, types.limit.lift_π_apply', fan.mk_π_app], exact h i j, }, { intros i j, convert congr_arg (limits.pi.π (λ p : ι × ι, F.obj (op (U p.1 ⊓ U p.2))) (i,j)) h, @@ -132,7 +132,7 @@ lemma is_gluing_iff_eq_res (sf : pi_opens F U) (s : F.obj (op (supr U))): begin split ; intros h, { ext ⟨i⟩, - rw [res, types.limit.lift_π_apply, fan.mk_π_app], + rw [res, types.limit.lift_π_apply', fan.mk_π_app], exact h i, }, { intro i, convert congr_arg (limits.pi.π (λ i : ι, F.obj (op (U i))) i) h, @@ -209,9 +209,8 @@ section local attribute [instance] concrete_category.has_coe_to_sort concrete_category.has_coe_to_fun -variables [has_limits_of_size.{w w} C] [reflects_isomorphisms (forget C)] - [preserves_limits_of_size.{w w} (forget C)] -variables {X : Top.{w}} (F : presheaf C X) {ι : Type w} (U : ι → opens X) +variables [has_limits C] [reflects_isomorphisms (forget C)] [preserves_limits (forget C)] +variables {X : Top.{v}} (F : presheaf C X) {ι : Type v} (U : ι → opens X) /-- For presheaves valued in a concrete category, whose forgetful functor reflects isomorphisms and @@ -236,10 +235,10 @@ section local attribute [instance] concrete_category.has_coe_to_sort concrete_category.has_coe_to_fun -variables [has_limits_of_size.{w w} C] [reflects_isomorphisms (concrete_category.forget C)] -variables [preserves_limits_of_size.{w w} (concrete_category.forget C)] +variables [has_limits C] [reflects_isomorphisms (concrete_category.forget C)] +variables [preserves_limits (concrete_category.forget C)] -variables {X : Top.{w}} (F : sheaf C X) {ι : Type w} (U : ι → opens X) +variables {X : Top.{v}} (F : sheaf C X) {ι : Type v} (U : ι → opens X) /-- A more convenient way of obtaining a unique gluing of sections for a sheaf. diff --git a/src/topology/sheaves/sheaf_of_functions.lean b/src/topology/sheaves/sheaf_of_functions.lean index 1ffac5ae6b1d9..8b3e8bcc0e660 100644 --- a/src/topology/sheaves/sheaf_of_functions.lean +++ b/src/topology/sheaves/sheaf_of_functions.lean @@ -33,11 +33,11 @@ open category_theory.limits open topological_space open topological_space.opens -universes v u +universe u noncomputable theory -variables (X : Top.{v}) +variables (X : Top.{u}) open Top @@ -101,13 +101,13 @@ namespace Top The sheaf of not-necessarily-continuous functions on `X` with values in type family `T : X → Type u`. -/ -def sheaf_to_Types (T : X → Type u) : sheaf (Type max v u) X := +def sheaf_to_Types (T : X → Type u) : sheaf (Type u) X := ⟨presheaf_to_Types X T, presheaf.to_Types_is_sheaf _ _⟩ /-- The sheaf of not-necessarily-continuous functions on `X` with values in a type `T`. -/ -def sheaf_to_Type (T : Type u) : sheaf (Type max v u) X := +def sheaf_to_Type (T : Type u) : sheaf (Type u) X := ⟨presheaf_to_Type X T, presheaf.to_Type_is_sheaf _ _⟩ end Top diff --git a/src/topology/sheaves/stalks.lean b/src/topology/sheaves/stalks.lean index c8564219b13bb..ea3a7426b187a 100644 --- a/src/topology/sheaves/stalks.lean +++ b/src/topology/sheaves/stalks.lean @@ -419,7 +419,7 @@ begin choose V m i₁ i₂ heq using λ x : U, F.presheaf.germ_eq x.1 x.2 x.2 s t (h x), -- Since `F` is a sheaf, we can prove the equality locally, if we can show that these -- neighborhoods form a cover of `U`. - apply Top.sheaf.eq_of_locally_eq'.{u v v} F V U i₁, + apply F.eq_of_locally_eq' V U i₁, { intros x hxU, rw [opens.mem_supr], exact ⟨⟨x, hxU⟩, m ⟨x, hxU⟩⟩ }, @@ -488,9 +488,9 @@ begin rw [opens.mem_supr], exact ⟨⟨x, hxU⟩, mV ⟨x, hxU⟩⟩ }, -- Since `F` is a sheaf, we can glue all the local preimages together to get a global preimage. - obtain ⟨s, s_spec, -⟩ := Top.sheaf.exists_unique_gluing'.{u v v} F V U iVU V_cover sf _, + obtain ⟨s, s_spec, -⟩ := F.exists_unique_gluing' V U iVU V_cover sf _, { use s, - apply Top.sheaf.eq_of_locally_eq'.{u v v} G V U iVU V_cover, + apply G.eq_of_locally_eq' V U iVU V_cover, intro x, rw [← comp_apply, ← f.1.naturality, comp_apply, s_spec, heq] }, { intros x y, From 8905e5ed90859939681a725b00f6063e65096d95 Mon Sep 17 00:00:00 2001 From: Anatole Dedecker Date: Thu, 6 Jul 2023 13:05:39 +0000 Subject: [PATCH 1236/1291] feat(topology/algebra/module/strong_topology): golf `arrow_congrSL` introduced in #19107 (#19128) I added more general definitions `precomp` and `postcomp` for expressing that (pre/post)-composing by a *fixed* continuous linear maps is continuous. These were planned about a year ago when I defined the strong topology and follow from [uniform_on_fun.precomp_uniform_continuous](https://leanprover-community.github.io/mathlib_docs/topology/uniform_space/uniform_convergence_topology.html#uniform_on_fun.precomp_uniform_continuous) and [uniform_on_fun.postcomp_uniform_continuous](https://leanprover-community.github.io/mathlib_docs/topology/uniform_space/uniform_convergence_topology.html#uniform_on_fun.postcomp_uniform_continuous). The proof of continuity of `arrow_congrSL` is a direct consequence of these, so we don't have to do it by hand. This is not really a "golf" since I added more lines than I removed, but these more general constructions will be needed at some point anyway (my use case was distribution theory) so I'm doing some proactive golfing :smile:. --- src/analysis/convolution.lean | 4 +- src/geometry/manifold/vector_bundle/hom.lean | 2 +- .../algebra/module/strong_topology.lean | 88 ++++++++++++------- src/topology/vector_bundle/hom.lean | 8 +- 4 files changed, 62 insertions(+), 40 deletions(-) diff --git a/src/analysis/convolution.lean b/src/analysis/convolution.lean index 74b30a8db71c0..199a713352ea1 100644 --- a/src/analysis/convolution.lean +++ b/src/analysis/convolution.lean @@ -1500,9 +1500,7 @@ begin simp only [ef, eg, comp_app, continuous_linear_equiv.apply_symm_apply, coe_comp', continuous_linear_equiv.prod_apply, continuous_linear_equiv.map_sub, continuous_linear_equiv.arrow_congr, continuous_linear_equiv.arrow_congrSL_symm_apply, - continuous_linear_equiv.coe_coe, comp_app, continuous_linear_equiv.apply_symm_apply, - linear_equiv.inv_fun_eq_symm, continuous_linear_equiv.arrow_congrₛₗ_symm_apply, - eq_self_iff_true] }, + continuous_linear_equiv.coe_coe, comp_app, continuous_linear_equiv.apply_symm_apply ] }, simp_rw [this] at A, exact A, end diff --git a/src/geometry/manifold/vector_bundle/hom.lean b/src/geometry/manifold/vector_bundle/hom.lean index e83f070f9d9a6..9dff42177cf64 100644 --- a/src/geometry/manifold/vector_bundle/hom.lean +++ b/src/geometry/manifold/vector_bundle/hom.lean @@ -70,7 +70,7 @@ begin simp only [continuous_linear_map_coord_change, continuous_linear_equiv.coe_coe, continuous_linear_equiv.arrow_congrSL_apply, comp_apply, function.comp, compL_apply, flip_apply, continuous_linear_equiv.symm_symm, linear_equiv.to_fun_eq_coe, - continuous_linear_equiv.arrow_congrₛₗ_apply, continuous_linear_map.coe_comp'] }, + continuous_linear_map.coe_comp'] }, end include _i₁ _i₂ diff --git a/src/topology/algebra/module/strong_topology.lean b/src/topology/algebra/module/strong_topology.lean index 706f6818c70fe..78e2334d92150 100644 --- a/src/topology/algebra/module/strong_topology.lean +++ b/src/topology/algebra/module/strong_topology.lean @@ -175,9 +175,12 @@ end general section bounded_sets -variables {𝕜₁ 𝕜₂ : Type*} [normed_field 𝕜₁] [normed_field 𝕜₂] {σ : 𝕜₁ →+* 𝕜₂} {E E' F F' : Type*} +variables {𝕜₁ 𝕜₂ 𝕜₃ : Type*} [normed_field 𝕜₁] [normed_field 𝕜₂] [normed_field 𝕜₃] + {σ : 𝕜₁ →+* 𝕜₂} {τ : 𝕜₂ →+* 𝕜₃} {ρ : 𝕜₁ →+* 𝕜₃} [ring_hom_comp_triple σ τ ρ] + {E E' F F' G : Type*} [add_comm_group E] [module 𝕜₁ E] [add_comm_group E'] [module ℝ E'] [add_comm_group F] [module 𝕜₂ F] [add_comm_group F'] [module ℝ F'] + [add_comm_group G] [module 𝕜₃ G] [topological_space E] /-- The topology of bounded convergence on `E →L[𝕜] F`. This coincides with the topology induced by @@ -224,6 +227,49 @@ protected lemma has_basis_nhds_zero [topological_space F] (λ SV, {f : E →SL[σ] F | ∀ x ∈ SV.1, f x ∈ SV.2}) := continuous_linear_map.has_basis_nhds_zero_of_basis (𝓝 0).basis_sets +variables (G) [topological_space F] [topological_space G] + +/-- Pre-composition by a *fixed* continuous linear map as a continuous linear map. +Note that in non-normed space it is not always true that composition is continuous +in both variables, so we have to fix one of them. -/ +@[simps] def precomp [topological_add_group G] [has_continuous_const_smul 𝕜₃ G] + [ring_hom_surjective σ] [ring_hom_isometric σ] (L : E →SL[σ] F) : + (F →SL[τ] G) →L[𝕜₃] (E →SL[ρ] G) := +{ to_fun := λ f, f.comp L, + map_add' := λ f g, add_comp f g L, + map_smul' := λ a f, smul_comp a f L, + cont := + begin + letI : uniform_space G := topological_add_group.to_uniform_space G, + haveI : uniform_add_group G := topological_add_comm_group_is_uniform, + rw (strong_topology.embedding_coe_fn _ _ _).continuous_iff, + refine (uniform_on_fun.precomp_uniform_continuous _).continuous.comp + (strong_topology.embedding_coe_fn _ _ _).continuous, + exact λ S hS, hS.image L, + end } + +variables (E) {G} + +/-- Post-composition by a *fixed* continuous linear map as a continuous linear map. +Note that in non-normed space it is not always true that composition is continuous +in both variables, so we have to fix one of them. -/ +@[simps] def postcomp [topological_add_group F] [topological_add_group G] + [has_continuous_const_smul 𝕜₃ G] [has_continuous_const_smul 𝕜₂ F] (L : F →SL[τ] G) : + (E →SL[σ] F) →SL[τ] (E →SL[ρ] G) := +{ to_fun := λ f, L.comp f, + map_add' := comp_add L, + map_smul' := comp_smulₛₗ L, + cont := + begin + letI : uniform_space G := topological_add_group.to_uniform_space G, + haveI : uniform_add_group G := topological_add_comm_group_is_uniform, + letI : uniform_space F := topological_add_group.to_uniform_space F, + haveI : uniform_add_group F := topological_add_comm_group_is_uniform, + rw (strong_topology.embedding_coe_fn _ _ _).continuous_iff, + exact (uniform_on_fun.postcomp_uniform_continuous L.uniform_continuous).continuous.comp + (strong_topology.embedding_coe_fn _ _ _).continuous + end } + end bounded_sets end continuous_linear_map @@ -249,49 +295,29 @@ variables {𝕜 : Type*} {𝕜₂ : Type*} {𝕜₃ : Type*} {𝕜₄ : Type*} [ring_hom_inv_pair σ₄₃ σ₃₄] [ring_hom_comp_triple σ₂₁ σ₁₄ σ₂₄] [ring_hom_comp_triple σ₂₄ σ₄₃ σ₂₃] [ring_hom_comp_triple σ₁₂ σ₂₃ σ₁₃] [ring_hom_comp_triple σ₁₃ σ₃₄ σ₁₄] + [ring_hom_comp_triple σ₂₃ σ₃₄ σ₂₄] [ring_hom_comp_triple σ₁₂ σ₂₄ σ₁₄] + [ring_hom_isometric σ₁₂] [ring_hom_isometric σ₂₁] include σ₁₄ σ₂₄ σ₁₃ σ₃₄ σ₂₁ σ₂₃ /-- A pair of continuous (semi)linear equivalences generates a (semi)linear equivalence between the spaces of continuous (semi)linear maps. -/ -@[simps] def arrow_congrₛₗ (e₁₂ : E ≃SL[σ₁₂] F) (e₄₃ : H ≃SL[σ₄₃] G) : - (E →SL[σ₁₄] H) ≃ₛₗ[σ₄₃] (F →SL[σ₂₃] G) := +@[simps] def arrow_congrSL (e₁₂ : E ≃SL[σ₁₂] F) (e₄₃ : H ≃SL[σ₄₃] G) : + (E →SL[σ₁₄] H) ≃SL[σ₄₃] (F →SL[σ₂₃] G) := { -- given explicitly to help `simps` to_fun := λ L, (e₄₃ : H →SL[σ₄₃] G).comp (L.comp (e₁₂.symm : F →SL[σ₂₁] E)), -- given explicitly to help `simps` inv_fun := λ L, (e₄₃.symm : G →SL[σ₃₄] H).comp (L.comp (e₁₂ : E →SL[σ₁₂] F)), map_add' := λ f g, by rw [add_comp, comp_add], map_smul' := λ t f, by rw [smul_comp, comp_smulₛₗ], + continuous_to_fun := + ((postcomp F e₄₃.to_continuous_linear_map).comp + (precomp H e₁₂.symm.to_continuous_linear_map)).continuous, + continuous_inv_fun := + ((precomp H e₁₂.to_continuous_linear_map).comp + (postcomp F e₄₃.symm.to_continuous_linear_map)).continuous, .. e₁₂.arrow_congr_equiv e₄₃, } -variables [ring_hom_isometric σ₂₁] - -lemma arrow_congrₛₗ_continuous (e₁₂ : E ≃SL[σ₁₂] F) (e₄₃ : H ≃SL[σ₄₃] G) : - continuous (id (e₁₂.arrow_congrₛₗ e₄₃ : (E →SL[σ₁₄] H) ≃ₛₗ[σ₄₃] (F →SL[σ₂₃] G))) := -begin - apply continuous_of_continuous_at_zero, - show filter.tendsto _ _ _, - simp_rw [(e₁₂.arrow_congrₛₗ e₄₃).map_zero], - rw continuous_linear_map.has_basis_nhds_zero.tendsto_iff - continuous_linear_map.has_basis_nhds_zero, - rintros ⟨sF, sG⟩ ⟨h1 : bornology.is_vonN_bounded 𝕜₂ sF, h2 : sG ∈ nhds (0:G)⟩, - dsimp, - refine ⟨(e₁₂.symm '' sF, e₄₃ ⁻¹' sG), ⟨h1.image (e₁₂.symm : F →SL[σ₂₁] E), _⟩, - λ _ h _ hx, h _ (set.mem_image_of_mem _ hx)⟩, - apply e₄₃.continuous.continuous_at, - simpa using h2, -end - -variables [ring_hom_isometric σ₁₂] - -/-- A pair of continuous (semi)linear equivalences generates an continuous (semi)linear equivalence -between the spaces of continuous (semi)linear maps. -/ -@[simps] def arrow_congrSL (e₁₂ : E ≃SL[σ₁₂] F) (e₄₃ : H ≃SL[σ₄₃] G) : - (E →SL[σ₁₄] H) ≃SL[σ₄₃] (F →SL[σ₂₃] G) := -{ continuous_to_fun := e₁₂.arrow_congrₛₗ_continuous e₄₃, - continuous_inv_fun := e₁₂.symm.arrow_congrₛₗ_continuous e₄₃.symm, - .. e₁₂.arrow_congrₛₗ e₄₃, } - end semilinear section linear diff --git a/src/topology/vector_bundle/hom.lean b/src/topology/vector_bundle/hom.lean index 7e09babbf0ef9..3ebe5d8e1e42c 100644 --- a/src/topology/vector_bundle/hom.lean +++ b/src/topology/vector_bundle/hom.lean @@ -110,9 +110,8 @@ begin { mfld_set_tac }, { intros b hb, ext L v, simp only [continuous_linear_map_coord_change, continuous_linear_equiv.coe_coe, - continuous_linear_equiv.arrow_congrₛₗ_apply, linear_equiv.to_fun_eq_coe, coe_comp', - continuous_linear_equiv.arrow_congrSL_apply, comp_apply, function.comp, compSL_apply, - flip_apply, continuous_linear_equiv.symm_symm] }, + continuous_linear_equiv.arrow_congrSL_apply, + comp_apply, function.comp, compSL_apply, flip_apply, continuous_linear_equiv.symm_symm] }, end omit iσ @@ -206,8 +205,7 @@ lemma continuous_linear_map_coord_change_apply (b : B) begin ext v, simp_rw [continuous_linear_map_coord_change, continuous_linear_equiv.coe_coe, - continuous_linear_equiv.arrow_congrSL_apply, linear_equiv.to_fun_eq_coe, - continuous_linear_equiv.arrow_congrₛₗ_apply, + continuous_linear_equiv.arrow_congrSL_apply, continuous_linear_map_apply, continuous_linear_map_symm_apply' σ e₁ e₂ hb.1, comp_apply, continuous_linear_equiv.coe_coe, continuous_linear_equiv.symm_symm, trivialization.continuous_linear_map_at_apply, trivialization.symmL_apply], From 1a51edf13debfcbe223fa06b1cb353b9ed9751cc Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Sat, 8 Jul 2023 12:34:58 +0000 Subject: [PATCH 1237/1291] chore(*): add mathlib4 synchronization comments (#19229) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.category.Algebra.basic` * `algebra.category.Algebra.limits` * `algebra.category.Module.change_of_rings` * `algebra.module.pid` * `algebraic_geometry.elliptic_curve.point` * `algebraic_geometry.morphisms.quasi_compact` * `algebraic_geometry.morphisms.quasi_separated` * `algebraic_geometry.projective_spectrum.structure_sheaf` * `algebraic_topology.fundamental_groupoid.induced_maps` * `algebraic_topology.fundamental_groupoid.product` * `algebraic_topology.fundamental_groupoid.simply_connected` * `analysis.normed.group.SemiNormedGroup.kernels` * `category_theory.adjunction.lifting` * `category_theory.monoidal.internal.Module` * `data.qpf.multivariate.constructions.cofix` * `geometry.manifold.diffeomorph` * `geometry.manifold.vector_bundle.smooth_section` * `geometry.manifold.whitney_embedding` * `group_theory.finite_abelian` * `linear_algebra.clifford_algebra.contraction` * `linear_algebra.exterior_algebra.grading` * `linear_algebra.matrix.schur_complement` * `linear_algebra.tensor_algebra.to_tensor_power` * `model_theory.direct_limit` * `model_theory.fraisse` * `number_theory.modular_forms.basic` * `representation_theory.character` * `representation_theory.group_cohomology.basic` * `ring_theory.jacobson` * `ring_theory.nullstellensatz` * `topology.metric_space.gromov_hausdorff` * `topology.vector_bundle.hom` --- src/algebra/category/Algebra/basic.lean | 3 +++ src/algebra/category/Algebra/limits.lean | 3 +++ src/algebra/category/Module/change_of_rings.lean | 3 +++ src/algebra/module/pid.lean | 3 +++ src/algebraic_geometry/elliptic_curve/point.lean | 3 +++ src/algebraic_geometry/morphisms/quasi_compact.lean | 3 +++ src/algebraic_geometry/morphisms/quasi_separated.lean | 3 +++ .../projective_spectrum/structure_sheaf.lean | 3 +++ src/algebraic_topology/fundamental_groupoid/induced_maps.lean | 3 +++ src/algebraic_topology/fundamental_groupoid/product.lean | 3 +++ .../fundamental_groupoid/simply_connected.lean | 3 +++ src/analysis/normed/group/SemiNormedGroup/kernels.lean | 3 +++ src/category_theory/adjunction/lifting.lean | 3 +++ src/category_theory/monoidal/internal/Module.lean | 3 +++ src/data/qpf/multivariate/constructions/cofix.lean | 3 +++ src/geometry/manifold/diffeomorph.lean | 3 +++ src/geometry/manifold/vector_bundle/smooth_section.lean | 3 +++ src/geometry/manifold/whitney_embedding.lean | 3 +++ src/group_theory/finite_abelian.lean | 3 +++ src/linear_algebra/clifford_algebra/contraction.lean | 3 +++ src/linear_algebra/exterior_algebra/grading.lean | 3 +++ src/linear_algebra/matrix/schur_complement.lean | 3 +++ src/linear_algebra/tensor_algebra/to_tensor_power.lean | 3 +++ src/model_theory/direct_limit.lean | 3 +++ src/model_theory/fraisse.lean | 3 +++ src/number_theory/modular_forms/basic.lean | 3 +++ src/representation_theory/character.lean | 3 +++ src/representation_theory/group_cohomology/basic.lean | 3 +++ src/ring_theory/jacobson.lean | 3 +++ src/ring_theory/nullstellensatz.lean | 3 +++ src/topology/metric_space/gromov_hausdorff.lean | 3 +++ src/topology/vector_bundle/hom.lean | 3 +++ 32 files changed, 96 insertions(+) diff --git a/src/algebra/category/Algebra/basic.lean b/src/algebra/category/Algebra/basic.lean index b0630ea5f9992..f8162553f0040 100644 --- a/src/algebra/category/Algebra/basic.lean +++ b/src/algebra/category/Algebra/basic.lean @@ -11,6 +11,9 @@ import algebra.category.Module.basic /-! # Category instance for algebras over a commutative ring +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We introduce the bundled category `Algebra` of algebras over a fixed commutative ring `R ` along with the forgetful functors to `Ring` and `Module`. We furthermore show that the functor associating to a type the free `R`-algebra on that type is left adjoint to the forgetful functor. diff --git a/src/algebra/category/Algebra/limits.lean b/src/algebra/category/Algebra/limits.lean index 19c423b60363e..abcf7968145a5 100644 --- a/src/algebra/category/Algebra/limits.lean +++ b/src/algebra/category/Algebra/limits.lean @@ -10,6 +10,9 @@ import algebra.category.Ring.limits /-! # The category of R-algebras has all limits +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Further, these limits are preserved by the forgetful functor --- that is, the underlying types are just the limits in the category of types. -/ diff --git a/src/algebra/category/Module/change_of_rings.lean b/src/algebra/category/Module/change_of_rings.lean index b52900dc79917..327a5e779e98a 100644 --- a/src/algebra/category/Module/change_of_rings.lean +++ b/src/algebra/category/Module/change_of_rings.lean @@ -9,6 +9,9 @@ import ring_theory.tensor_product /-! # Change Of Rings +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions * `category_theory.Module.restrict_scalars`: given rings `R, S` and a ring homomorphism `R ⟶ S`, diff --git a/src/algebra/module/pid.lean b/src/algebra/module/pid.lean index fe80596510b4a..adc0435a263cd 100644 --- a/src/algebra/module/pid.lean +++ b/src/algebra/module/pid.lean @@ -11,6 +11,9 @@ import algebra.category.Module.biproducts /-! # Structure of finitely generated modules over a PID +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main statements * `module.equiv_direct_sum_of_is_torsion` : A finitely generated torsion module over a PID is diff --git a/src/algebraic_geometry/elliptic_curve/point.lean b/src/algebraic_geometry/elliptic_curve/point.lean index e347de9ad650d..f8e6785033fa0 100644 --- a/src/algebraic_geometry/elliptic_curve/point.lean +++ b/src/algebraic_geometry/elliptic_curve/point.lean @@ -11,6 +11,9 @@ import ring_theory.class_group /-! # Nonsingular rational points on Weierstrass curves +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the type of nonsingular rational points on a Weierstrass curve over a field and proves that it forms an abelian group under a geometric secant-and-tangent process. diff --git a/src/algebraic_geometry/morphisms/quasi_compact.lean b/src/algebraic_geometry/morphisms/quasi_compact.lean index b3281898a3059..b0914f8516fd2 100644 --- a/src/algebraic_geometry/morphisms/quasi_compact.lean +++ b/src/algebraic_geometry/morphisms/quasi_compact.lean @@ -10,6 +10,9 @@ import algebraic_geometry.limits /-! # Quasi-compact morphisms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A morphism of schemes is quasi-compact if the preimages of quasi-compact open sets are quasi-compact. diff --git a/src/algebraic_geometry/morphisms/quasi_separated.lean b/src/algebraic_geometry/morphisms/quasi_separated.lean index 8b75d385dcf82..33569b42b09e2 100644 --- a/src/algebraic_geometry/morphisms/quasi_separated.lean +++ b/src/algebraic_geometry/morphisms/quasi_separated.lean @@ -9,6 +9,9 @@ import topology.quasi_separated /-! # Quasi-separated morphisms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A morphism of schemes `f : X ⟶ Y` is quasi-separated if the diagonal morphism `X ⟶ X ×[Y] X` is quasi-compact. diff --git a/src/algebraic_geometry/projective_spectrum/structure_sheaf.lean b/src/algebraic_geometry/projective_spectrum/structure_sheaf.lean index d6d743bec77d1..9003dcd0fa7d1 100644 --- a/src/algebraic_geometry/projective_spectrum/structure_sheaf.lean +++ b/src/algebraic_geometry/projective_spectrum/structure_sheaf.lean @@ -11,6 +11,9 @@ import algebraic_geometry.locally_ringed_space /-! # The structure sheaf on `projective_spectrum 𝒜`. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In `src/algebraic_geometry/topology.lean`, we have given a topology on `projective_spectrum 𝒜`; in this file we will construct a sheaf on `projective_spectrum 𝒜`. diff --git a/src/algebraic_topology/fundamental_groupoid/induced_maps.lean b/src/algebraic_topology/fundamental_groupoid/induced_maps.lean index a1c1a2b4c194c..1536082abcf65 100644 --- a/src/algebraic_topology/fundamental_groupoid/induced_maps.lean +++ b/src/algebraic_topology/fundamental_groupoid/induced_maps.lean @@ -10,6 +10,9 @@ import algebraic_topology.fundamental_groupoid.product /-! # Homotopic maps induce naturally isomorphic functors +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions - `fundamental_groupoid_functor.homotopic_maps_nat_iso H` The natural isomorphism diff --git a/src/algebraic_topology/fundamental_groupoid/product.lean b/src/algebraic_topology/fundamental_groupoid/product.lean index 8ecf0078ff45d..7d8a7c033710a 100644 --- a/src/algebraic_topology/fundamental_groupoid/product.lean +++ b/src/algebraic_topology/fundamental_groupoid/product.lean @@ -11,6 +11,9 @@ import topology.homotopy.product /-! # Fundamental groupoid preserves products + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. In this file, we give the following definitions/theorems: - `fundamental_groupoid_functor.pi_iso` An isomorphism between Π i, (π Xᵢ) and π (Πi, Xᵢ), whose diff --git a/src/algebraic_topology/fundamental_groupoid/simply_connected.lean b/src/algebraic_topology/fundamental_groupoid/simply_connected.lean index e04bbe10c63b6..05e4e9f523fa4 100644 --- a/src/algebraic_topology/fundamental_groupoid/simply_connected.lean +++ b/src/algebraic_topology/fundamental_groupoid/simply_connected.lean @@ -10,6 +10,9 @@ import algebraic_topology.fundamental_groupoid.punit /-! # Simply connected spaces + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. This file defines simply connected spaces. A topological space is simply connected if its fundamental groupoid is equivalent to `unit`. diff --git a/src/analysis/normed/group/SemiNormedGroup/kernels.lean b/src/analysis/normed/group/SemiNormedGroup/kernels.lean index 0dc24df7e08d0..b1822fbc5011c 100644 --- a/src/analysis/normed/group/SemiNormedGroup/kernels.lean +++ b/src/analysis/normed/group/SemiNormedGroup/kernels.lean @@ -10,6 +10,9 @@ import category_theory.limits.shapes.kernels /-! # Kernels and cokernels in SemiNormedGroup₁ and SemiNormedGroup +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We show that `SemiNormedGroup₁` has cokernels (for which of course the `cokernel.π f` maps are norm non-increasing), as well as the easier result that `SemiNormedGroup` has cokernels. We also show that diff --git a/src/category_theory/adjunction/lifting.lean b/src/category_theory/adjunction/lifting.lean index f8fa2872d8234..d1b70e8462e1a 100644 --- a/src/category_theory/adjunction/lifting.lean +++ b/src/category_theory/adjunction/lifting.lean @@ -11,6 +11,9 @@ import category_theory.monad.coequalizer /-! # Adjoint lifting +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file gives two constructions for building left adjoints: the adjoint triangle theorem and the adjoint lifting theorem. The adjoint triangle theorem says that given a functor `U : B ⥤ C` with a left adjoint `F` such diff --git a/src/category_theory/monoidal/internal/Module.lean b/src/category_theory/monoidal/internal/Module.lean index 27071854b7d83..a8dcbc32b37ac 100644 --- a/src/category_theory/monoidal/internal/Module.lean +++ b/src/category_theory/monoidal/internal/Module.lean @@ -11,6 +11,9 @@ import category_theory.monoidal.Mon_ /-! # `Mon_ (Module R) ≌ Algebra R` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The category of internal monoid objects in `Module R` is equivalent to the category of "native" bundled `R`-algebras. diff --git a/src/data/qpf/multivariate/constructions/cofix.lean b/src/data/qpf/multivariate/constructions/cofix.lean index ba67e60e0df8b..93bc6f9098c1e 100644 --- a/src/data/qpf/multivariate/constructions/cofix.lean +++ b/src/data/qpf/multivariate/constructions/cofix.lean @@ -12,6 +12,9 @@ import data.qpf.multivariate.basic /-! # The final co-algebra of a multivariate qpf is again a qpf. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + For a `(n+1)`-ary QPF `F (α₀,..,αₙ)`, we take the least fixed point of `F` with regards to its last argument `αₙ`. The result is a `n`-ary functor: `fix F (α₀,..,αₙ₋₁)`. Making `fix F` into a functor allows us to take the fixed point, compose with other functors diff --git a/src/geometry/manifold/diffeomorph.lean b/src/geometry/manifold/diffeomorph.lean index 4a1a820e6dfda..f7eef23e11ca7 100644 --- a/src/geometry/manifold/diffeomorph.lean +++ b/src/geometry/manifold/diffeomorph.lean @@ -9,6 +9,9 @@ import geometry.manifold.cont_mdiff_mfderiv /-! # Diffeomorphisms + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. This file implements diffeomorphisms. ## Definitions diff --git a/src/geometry/manifold/vector_bundle/smooth_section.lean b/src/geometry/manifold/vector_bundle/smooth_section.lean index c7782d18cc4b6..d6c5115c08f70 100644 --- a/src/geometry/manifold/vector_bundle/smooth_section.lean +++ b/src/geometry/manifold/vector_bundle/smooth_section.lean @@ -11,6 +11,9 @@ import geometry.manifold.algebra.lie_group /-! # Smooth sections +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define the type `cont_mdiff_section` of `n` times continuously differentiable sections of a smooth vector bundle over a manifold `M` and prove that it's a module. -/ diff --git a/src/geometry/manifold/whitney_embedding.lean b/src/geometry/manifold/whitney_embedding.lean index 736ba534e91f6..c47bcd133c47e 100644 --- a/src/geometry/manifold/whitney_embedding.lean +++ b/src/geometry/manifold/whitney_embedding.lean @@ -10,6 +10,9 @@ import geometry.manifold.partition_of_unity /-! # Whitney embedding theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove a version of the Whitney embedding theorem: for any compact real manifold `M`, for sufficiently large `n` there exists a smooth embedding `M → ℝ^n`. diff --git a/src/group_theory/finite_abelian.lean b/src/group_theory/finite_abelian.lean index 0a844cd0e78bc..bf23854627a70 100644 --- a/src/group_theory/finite_abelian.lean +++ b/src/group_theory/finite_abelian.lean @@ -9,6 +9,9 @@ import data.zmod.quotient /-! # Structure of finite(ly generated) abelian groups +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + * `add_comm_group.equiv_free_prod_direct_sum_zmod` : Any finitely generated abelian group is the product of a power of `ℤ` and a direct sum of some `zmod (p i ^ e i)` for some prime powers `p i ^ e i`. diff --git a/src/linear_algebra/clifford_algebra/contraction.lean b/src/linear_algebra/clifford_algebra/contraction.lean index 152ebc5bd63b3..6384379fe7097 100644 --- a/src/linear_algebra/clifford_algebra/contraction.lean +++ b/src/linear_algebra/clifford_algebra/contraction.lean @@ -11,6 +11,9 @@ import linear_algebra.clifford_algebra.conjugation /-! # Contraction in Clifford Algebras +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains some of the results from [grinberg_clifford_2016][]. The key result is `clifford_algebra.equiv_exterior`. diff --git a/src/linear_algebra/exterior_algebra/grading.lean b/src/linear_algebra/exterior_algebra/grading.lean index eee9bc304edb2..0f5089ccbee07 100644 --- a/src/linear_algebra/exterior_algebra/grading.lean +++ b/src/linear_algebra/exterior_algebra/grading.lean @@ -9,6 +9,9 @@ import ring_theory.graded_algebra.basic /-! # Results about the grading structure of the exterior algebra +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Many of these results are copied with minimal modification from the tensor algebra. The main result is `exterior_algebra.graded_algebra`, which says that the exterior algebra is a diff --git a/src/linear_algebra/matrix/schur_complement.lean b/src/linear_algebra/matrix/schur_complement.lean index 1c354e0f09131..4e3cc43ea1180 100644 --- a/src/linear_algebra/matrix/schur_complement.lean +++ b/src/linear_algebra/matrix/schur_complement.lean @@ -9,6 +9,9 @@ import linear_algebra.matrix.pos_def /-! # 2×2 block matrices and the Schur complement +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves properties of 2×2 block matrices `[A B; C D]` that relate to the Schur complement `D - C⬝A⁻¹⬝B`. diff --git a/src/linear_algebra/tensor_algebra/to_tensor_power.lean b/src/linear_algebra/tensor_algebra/to_tensor_power.lean index 5d619d69bba0b..d79a8e426472b 100644 --- a/src/linear_algebra/tensor_algebra/to_tensor_power.lean +++ b/src/linear_algebra/tensor_algebra/to_tensor_power.lean @@ -8,6 +8,9 @@ import linear_algebra.tensor_power /-! # Tensor algebras as direct sums of tensor powers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we show that `tensor_algebra R M` is isomorphic to a direct sum of tensor powers, as `tensor_algebra.equiv_direct_sum`. -/ diff --git a/src/model_theory/direct_limit.lean b/src/model_theory/direct_limit.lean index b228a91cfa896..7833595665f85 100644 --- a/src/model_theory/direct_limit.lean +++ b/src/model_theory/direct_limit.lean @@ -10,6 +10,9 @@ import model_theory.finitely_generated /-! # Direct Limits of First-Order Structures + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. This file constructs the direct limit of a directed system of first-order embeddings. ## Main Definitions diff --git a/src/model_theory/fraisse.lean b/src/model_theory/fraisse.lean index d819e48cea930..2f372ad156a27 100644 --- a/src/model_theory/fraisse.lean +++ b/src/model_theory/fraisse.lean @@ -11,6 +11,9 @@ import model_theory.bundled /-! # Fraïssé Classes and Fraïssé Limits +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file pertains to the ages of countable first-order structures. The age of a structure is the class of all finitely-generated structures that embed into it. diff --git a/src/number_theory/modular_forms/basic.lean b/src/number_theory/modular_forms/basic.lean index 1d1346c5dfee0..01c52760e61b2 100644 --- a/src/number_theory/modular_forms/basic.lean +++ b/src/number_theory/modular_forms/basic.lean @@ -9,6 +9,9 @@ import number_theory.modular_forms.slash_invariant_forms /-! # Modular forms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines modular forms and proves some basic properties about them. We begin by defining modular forms and cusp forms as extension of `slash_invariant_forms` then we diff --git a/src/representation_theory/character.lean b/src/representation_theory/character.lean index e346efe45599a..07ddd2fbf264d 100644 --- a/src/representation_theory/character.lean +++ b/src/representation_theory/character.lean @@ -10,6 +10,9 @@ import representation_theory.invariants /-! # Characters of representations +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file introduces characters of representation and proves basic lemmas about how characters behave under various operations on representations. diff --git a/src/representation_theory/group_cohomology/basic.lean b/src/representation_theory/group_cohomology/basic.lean index a972d9f61f849..b2860c0192415 100644 --- a/src/representation_theory/group_cohomology/basic.lean +++ b/src/representation_theory/group_cohomology/basic.lean @@ -10,6 +10,9 @@ import representation_theory.group_cohomology.resolution /-! # The group cohomology of a `k`-linear `G`-representation +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Let `k` be a commutative ring and `G` a group. This file defines the group cohomology of `A : Rep k G` to be the cohomology of the complex $$0 \to \mathrm{Fun}(G^0, A) \to \mathrm{Fun}(G^1, A) \to \mathrm{Fun}(G^2, A) \to \dots$$ diff --git a/src/ring_theory/jacobson.lean b/src/ring_theory/jacobson.lean index a44c12838c5c5..8c0a791739917 100644 --- a/src/ring_theory/jacobson.lean +++ b/src/ring_theory/jacobson.lean @@ -9,6 +9,9 @@ import ring_theory.jacobson_ideal /-! # Jacobson Rings + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. The following conditions are equivalent for a ring `R`: 1. Every radical ideal `I` is equal to its Jacobson radical 2. Every radical ideal `I` can be written as an intersection of maximal ideals diff --git a/src/ring_theory/nullstellensatz.lean b/src/ring_theory/nullstellensatz.lean index ab70ecadfb579..def2196661005 100644 --- a/src/ring_theory/nullstellensatz.lean +++ b/src/ring_theory/nullstellensatz.lean @@ -10,6 +10,9 @@ import algebraic_geometry.prime_spectrum.basic /-! # Nullstellensatz + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. This file establishes a version of Hilbert's classical Nullstellensatz for `mv_polynomial`s. The main statement of the theorem is `vanishing_ideal_zero_locus_eq_radical`. diff --git a/src/topology/metric_space/gromov_hausdorff.lean b/src/topology/metric_space/gromov_hausdorff.lean index d6b068acf48f9..f3d546bb9f5dc 100644 --- a/src/topology/metric_space/gromov_hausdorff.lean +++ b/src/topology/metric_space/gromov_hausdorff.lean @@ -12,6 +12,9 @@ import topology.metric_space.kuratowski /-! # Gromov-Hausdorff distance +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the Gromov-Hausdorff distance on the space of nonempty compact metric spaces up to isometry. diff --git a/src/topology/vector_bundle/hom.lean b/src/topology/vector_bundle/hom.lean index 3ebe5d8e1e42c..96ea682a0be2b 100644 --- a/src/topology/vector_bundle/hom.lean +++ b/src/topology/vector_bundle/hom.lean @@ -10,6 +10,9 @@ import analysis.normed_space.operator_norm /-! # The vector bundle of continuous (semi)linear maps +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We define the (topological) vector bundle of continuous (semi)linear maps between two vector bundles over the same base. From bd365b1a4901dbd878e86cb146c2bd86533df468 Mon Sep 17 00:00:00 2001 From: b-reinke Date: Tue, 11 Jul 2023 12:40:13 +0000 Subject: [PATCH 1238/1291] feat(group_theory/sylow): add inverse to card_eq_multiplicity (#18300) The lemma `card_eq_multiplicity` states that a Sylow group of a finite group has cardinality p^n, where n is the multiplicity of p in the group order. This PR adds an inverse definition `card_eq_multiplicity_to_sylow`, promoting a subgroup of the right cardinality to a Sylow group, and a simplification lemma `coe_card_eq_multiplicity_to_sylow` for the coercion of the resulting Sylow group back to a subgroup. --- src/group_theory/sylow.lean | 27 +++++++++++++++++++++------ 1 file changed, 21 insertions(+), 6 deletions(-) diff --git a/src/group_theory/sylow.lean b/src/group_theory/sylow.lean index 91d962c8014b9..8b32b29b66805 100644 --- a/src/group_theory/sylow.lean +++ b/src/group_theory/sylow.lean @@ -36,7 +36,7 @@ The Sylow theorems are the following results for every finite group `G` and ever there exists a subgroup of `G` of order `pⁿ`. * `is_p_group.exists_le_sylow`: A generalization of Sylow's first theorem: Every `p`-subgroup is contained in a Sylow `p`-subgroup. -* `sylow.card_eq_multiplicity`: The cardinality of a Sylow group is `p ^ n` +* `sylow.card_eq_multiplicity`: The cardinality of a Sylow subgroup is `p ^ n` where `n` is the multiplicity of `p` in the group order. * `sylow_conjugate`: A generalization of Sylow's second theorem: If the number of Sylow `p`-subgroups is finite, then all Sylow `p`-subgroups are conjugate. @@ -592,7 +592,7 @@ begin rwa [h, card_bot] at key, end -/-- The cardinality of a Sylow group is `p ^ n` +/-- The cardinality of a Sylow subgroup is `p ^ n` where `n` is the multiplicity of `p` in the group order. -/ lemma card_eq_multiplicity [fintype G] {p : ℕ} [hp : fact p.prime] (P : sylow p G) : card P = p ^ nat.factorization (card G) p := @@ -603,6 +603,21 @@ begin exact P.1.card_subgroup_dvd_card, end +/-- A subgroup with cardinality `p ^ n` is a Sylow subgroup + where `n` is the multiplicity of `p` in the group order. -/ +def of_card [fintype G] {p : ℕ} [hp : fact p.prime] (H : subgroup G) [fintype H] + (card_eq : card H = p ^ (card G).factorization p) : sylow p G := +{ to_subgroup := H, + is_p_group' := is_p_group.of_card card_eq, + is_maximal' := begin + obtain ⟨P, hHP⟩ := (is_p_group.of_card card_eq).exists_le_sylow, + exact set_like.ext' (set.eq_of_subset_of_card_le hHP + (P.card_eq_multiplicity.trans card_eq.symm).le).symm ▸ λ _, P.3, + end } + +@[simp, norm_cast] lemma coe_of_card [fintype G] {p : ℕ} [hp : fact p.prime] (H : subgroup G) + [fintype H] (card_eq : card H = p ^ (card G).factorization p) : ↑(of_card H card_eq) = H := rfl + lemma subsingleton_of_normal {p : ℕ} [fact p.prime] [finite (sylow p G)] (P : sylow p G) (h : (P : subgroup G).normal) : subsingleton (sylow p G) := begin @@ -667,8 +682,8 @@ normalizer_eq_top.mp $ normalizer_condition_iff_only_full_group_self_normalizing open_locale big_operators -/-- If all its sylow groups are normal, then a finite group is isomorphic to the direct product -of these sylow groups. +/-- If all its Sylow subgroups are normal, then a finite group is isomorphic to the direct product +of these Sylow subgroups. -/ noncomputable def direct_product_of_normal [fintype G] @@ -677,7 +692,7 @@ def direct_product_of_normal [fintype G] begin set ps := (fintype.card G).factorization.support, - -- “The” sylow group for p + -- “The” Sylow subgroup for p let P : Π p, sylow p G := default, have hcomm : pairwise (λ (p₁ p₂ : ps), ∀ (x y : G), x ∈ P p₁ → y ∈ P p₂ → commute x y), @@ -689,7 +704,7 @@ begin apply is_p_group.disjoint_of_ne p₁ p₂ hne' _ _ (P p₁).is_p_group' (P p₂).is_p_group', }, refine mul_equiv.trans _ _, - -- There is only one sylow group for each p, so the inner product is trivial + -- There is only one Sylow subgroup for each p, so the inner product is trivial show (Π p : ps, Π P : sylow p G, P) ≃* (Π p : ps, P p), { -- here we need to help the elaborator with an explicit instantiation apply @mul_equiv.Pi_congr_right ps (λ p, (Π P : sylow p G, P)) (λ p, P p) _ _ , From 3ed3f98a1e836241990d3d308f1577e434977130 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 11 Jul 2023 12:40:14 +0000 Subject: [PATCH 1239/1291] =?UTF-8?q?feat(data/nat/order/basic):=20`a=20+?= =?UTF-8?q?=20b=20-=201=20=E2=89=A4=20a=20*=20b`=20(#18737)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit and golf `max_eq_zero_iff`/`min_eq_zero_iff` To be used for Kneser's addition theorem. Co-authored-by: Yury G. Kudryashov --- src/data/nat/order/basic.lean | 32 ++++++++++---------------------- 1 file changed, 10 insertions(+), 22 deletions(-) diff --git a/src/data/nat/order/basic.lean b/src/data/nat/order/basic.lean index fcd6864736ece..c0913019067e2 100644 --- a/src/data/nat/order/basic.lean +++ b/src/data/nat/order/basic.lean @@ -94,29 +94,9 @@ lemma eq_zero_of_mul_le (hb : 2 ≤ n) (h : n * m ≤ m) : m = 0 := eq_zero_of_double_le $ le_trans (nat.mul_le_mul_right _ hb) h lemma zero_max : max 0 n = n := max_eq_right (zero_le _) -@[simp] lemma min_eq_zero_iff : min m n = 0 ↔ m = 0 ∨ n = 0 := -begin - split, - { intro h, - cases le_total m n with H H, - { simpa [H] using or.inl h }, - { simpa [H] using or.inr h } }, - { rintro (rfl|rfl); - simp } -end -@[simp] lemma max_eq_zero_iff : max m n = 0 ↔ m = 0 ∧ n = 0 := -begin - split, - { intro h, - cases le_total m n with H H, - { simp only [H, max_eq_right] at h, - exact ⟨le_antisymm (H.trans h.le) (zero_le _), h⟩ }, - { simp only [H, max_eq_left] at h, - exact ⟨h, le_antisymm (H.trans h.le) (zero_le _)⟩ } }, - { rintro ⟨rfl, rfl⟩, - simp } -end +@[simp] lemma min_eq_zero_iff : min m n = 0 ↔ m = 0 ∨ n = 0 := min_eq_bot +@[simp] lemma max_eq_zero_iff : max m n = 0 ↔ m = 0 ∧ n = 0 := max_eq_bot lemma add_eq_max_iff : m + n = max m n ↔ m = 0 ∨ n = 0 := begin @@ -286,6 +266,14 @@ end | 0 := iff_of_false (lt_irrefl _) zero_le_one.not_lt | (n + 1) := lt_mul_iff_one_lt_left n.succ_pos +lemma add_sub_one_le_mul (hm : m ≠ 0) (hn : n ≠ 0) : m + n - 1 ≤ m * n := +begin + cases m, + { cases hm rfl }, + { rw [succ_add, succ_sub_one, succ_mul], + exact add_le_add_right (le_mul_of_one_le_right' $ pos_iff_ne_zero.2 hn) _ } +end + /-! ### Recursion and induction principles From 4e24c4bfcff371c71f7ba22050308aa17815626c Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 11 Jul 2023 15:46:03 +0000 Subject: [PATCH 1240/1291] feat(data/set/intervals/proj_Icc): Extending from `Ici a` (#18795) Define the `Ici` and `Iic` versions of `set.proj_Icc`/`set.Icc_extend`. Slightly reorder the lemmas to allow better overall structure. This is useful for extending convex functions by a junk value. --- src/data/set/intervals/proj_Icc.lean | 152 ++++++++++++++++++++++----- 1 file changed, 127 insertions(+), 25 deletions(-) diff --git a/src/data/set/intervals/proj_Icc.lean b/src/data/set/intervals/proj_Icc.lean index e6cb3cff298e2..41a033fc0ed3a 100644 --- a/src/data/set/intervals/proj_Icc.lean +++ b/src/data/set/intervals/proj_Icc.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury G. Kudryashov, Patrick Massot -/ import data.set.function -import data.set.intervals.basic +import data.set.intervals.ord_connected /-! # Projection of a line onto a closed interval @@ -14,8 +14,16 @@ import data.set.intervals.basic Given a linearly ordered type `α`, in this file we define +* `set.proj_Ici (a : α)` to be the map `α → [a, ∞[` sending `]-∞, a]` to `a`, + and each point `x ∈ [a, ∞[` to itself; +* `set.proj_Iic (b : α)` to be the map `α → ]-∞, b[` sending `[b, ∞[` to `b`, + and each point `x ∈ ]-∞, b]` to itself; * `set.proj_Icc (a b : α) (h : a ≤ b)` to be the map `α → [a, b]` sending `(-∞, a]` to `a`, `[b, ∞)` to `b`, and each point `x ∈ [a, b]` to itself; +* `set.Ici_extend {a : α} (f : Ici a → β)` to be the extension of `f` to `α` defined + as `f ∘ proj_Ici a`. +* `set.Iic_extend {b : α} (f : Iic b → β)` to be the extension of `f` to `α` defined + as `f ∘ proj_Iic b`. * `set.Icc_extend {a b : α} (h : a ≤ b) (f : Icc a b → β)` to be the extension of `f` to `α` defined as `f ∘ proj_Icc a b h`. @@ -28,88 +36,156 @@ open function namespace set +/-- Projection of `α` to the closed interval `[a, ∞[`. -/ +def proj_Ici (a x : α) : Ici a := ⟨max a x, le_max_left _ _⟩ + +/-- Projection of `α` to the closed interval `]-∞, b]`. -/ +def proj_Iic (b x : α) : Iic b := ⟨min b x, min_le_left _ _⟩ + /-- Projection of `α` to the closed interval `[a, b]`. -/ def proj_Icc (a b : α) (h : a ≤ b) (x : α) : Icc a b := ⟨max a (min b x), le_max_left _ _, max_le h (min_le_left _ _)⟩ variables {a b : α} (h : a ≤ b) {x : α} +@[norm_cast] lemma coe_proj_Ici (a x : α) : (proj_Ici a x : α) = max a x := rfl +@[norm_cast] lemma coe_proj_Iic (b x : α) : (proj_Iic b x : α) = min b x := rfl +@[norm_cast] lemma coe_proj_Icc (a b : α) (h : a ≤ b) (x : α) : + (proj_Icc a b h x : α) = max a (min b x) := rfl + +lemma proj_Ici_of_le (hx : x ≤ a) : proj_Ici a x = ⟨a, le_rfl⟩ := subtype.ext $ max_eq_left hx +lemma proj_Iic_of_le (hx : b ≤ x) : proj_Iic b x = ⟨b, le_rfl⟩ := subtype.ext $ min_eq_left hx + lemma proj_Icc_of_le_left (hx : x ≤ a) : proj_Icc a b h x = ⟨a, left_mem_Icc.2 h⟩ := by simp [proj_Icc, hx, hx.trans h] -@[simp] lemma proj_Icc_left : proj_Icc a b h a = ⟨a, left_mem_Icc.2 h⟩ := -proj_Icc_of_le_left h le_rfl - lemma proj_Icc_of_right_le (hx : b ≤ x) : proj_Icc a b h x = ⟨b, right_mem_Icc.2 h⟩ := by simp [proj_Icc, hx, h] +@[simp] lemma proj_Ici_self (a : α) : proj_Ici a a = ⟨a, le_rfl⟩ := proj_Ici_of_le le_rfl +@[simp] lemma proj_Iic_self (b : α) : proj_Iic b b = ⟨b, le_rfl⟩ := proj_Iic_of_le le_rfl + +@[simp] lemma proj_Icc_left : proj_Icc a b h a = ⟨a, left_mem_Icc.2 h⟩ := +proj_Icc_of_le_left h le_rfl + @[simp] lemma proj_Icc_right : proj_Icc a b h b = ⟨b, right_mem_Icc.2 h⟩ := proj_Icc_of_right_le h le_rfl +lemma proj_Ici_eq_self : proj_Ici a x = ⟨a, le_rfl⟩ ↔ x ≤ a := by simp [proj_Ici, subtype.ext_iff] +lemma proj_Iic_eq_self : proj_Iic b x = ⟨b, le_rfl⟩ ↔ b ≤ x := by simp [proj_Iic, subtype.ext_iff] + lemma proj_Icc_eq_left (h : a < b) : proj_Icc a b h.le x = ⟨a, left_mem_Icc.mpr h.le⟩ ↔ x ≤ a := -begin - refine ⟨λ h', _, proj_Icc_of_le_left _⟩, - simp_rw [subtype.ext_iff_val, proj_Icc, max_eq_left_iff, min_le_iff, h.not_le, false_or] at h', - exact h' -end +by simp [proj_Icc, subtype.ext_iff, h.not_le] lemma proj_Icc_eq_right (h : a < b) : proj_Icc a b h.le x = ⟨b, right_mem_Icc.mpr h.le⟩ ↔ b ≤ x := -begin - refine ⟨λ h', _, proj_Icc_of_right_le _⟩, - simp_rw [subtype.ext_iff_val, proj_Icc] at h', - have := ((max_choice _ _).resolve_left (by simp [h.ne', h'])).symm.trans h', - exact min_eq_left_iff.mp this -end +by simp [proj_Icc, subtype.ext_iff, max_min_distrib_left, h.le, h.not_le] +lemma proj_Ici_of_mem (hx : x ∈ Ici a) : proj_Ici a x = ⟨x, hx⟩ := by simpa [proj_Ici] +lemma proj_Iic_of_mem (hx : x ∈ Iic b) : proj_Iic b x = ⟨x, hx⟩ := by simpa [proj_Iic] lemma proj_Icc_of_mem (hx : x ∈ Icc a b) : proj_Icc a b h x = ⟨x, hx⟩ := by simp [proj_Icc, hx.1, hx.2] +@[simp] lemma proj_Ici_coe (x : Ici a) : proj_Ici a x = x := by { cases x, apply proj_Ici_of_mem } +@[simp] lemma proj_Iic_coe (x : Iic b) : proj_Iic b x = x := by { cases x, apply proj_Iic_of_mem } @[simp] lemma proj_Icc_coe (x : Icc a b) : proj_Icc a b h x = x := by { cases x, apply proj_Icc_of_mem } +lemma proj_Ici_surj_on : surj_on (proj_Ici a) (Ici a) univ := λ x _, ⟨x, x.2, proj_Ici_coe x⟩ +lemma proj_Iic_surj_on : surj_on (proj_Iic b) (Iic b) univ := λ x _, ⟨x, x.2, proj_Iic_coe x⟩ lemma proj_Icc_surj_on : surj_on (proj_Icc a b h) (Icc a b) univ := λ x _, ⟨x, x.2, proj_Icc_coe h x⟩ -lemma proj_Icc_surjective : surjective (proj_Icc a b h) := -λ x, ⟨x, proj_Icc_coe h x⟩ +lemma proj_Ici_surjective : surjective (proj_Ici a) := λ x, ⟨x, proj_Ici_coe x⟩ +lemma proj_Iic_surjective : surjective (proj_Iic b) := λ x, ⟨x, proj_Iic_coe x⟩ +lemma proj_Icc_surjective : surjective (proj_Icc a b h) := λ x, ⟨x, proj_Icc_coe h x⟩ -@[simp] lemma range_proj_Icc : range (proj_Icc a b h) = univ := -(proj_Icc_surjective h).range_eq +@[simp] lemma range_proj_Ici : range (proj_Ici a) = univ := proj_Ici_surjective.range_eq +@[simp] lemma range_proj_Iic : range (proj_Iic a) = univ := proj_Iic_surjective.range_eq +@[simp] lemma range_proj_Icc : range (proj_Icc a b h) = univ := (proj_Icc_surjective h).range_eq +lemma monotone_proj_Ici : monotone (proj_Ici a) := λ x y, max_le_max le_rfl +lemma monotone_proj_Iic : monotone (proj_Iic a) := λ x y, min_le_min le_rfl lemma monotone_proj_Icc : monotone (proj_Icc a b h) := λ x y hxy, max_le_max le_rfl $ min_le_min le_rfl hxy +lemma strict_mono_on_proj_Ici : strict_mono_on (proj_Ici a) (Ici a) := +λ x hx y hy hxy, by simpa only [proj_Ici_of_mem, hx, hy] +lemma strict_mono_on_proj_Iic : strict_mono_on (proj_Iic b) (Iic b) := +λ x hx y hy hxy, by simpa only [proj_Iic_of_mem, hx, hy] lemma strict_mono_on_proj_Icc : strict_mono_on (proj_Icc a b h) (Icc a b) := λ x hx y hy hxy, by simpa only [proj_Icc_of_mem, hx, hy] +/-- Extend a function `[a, ∞[ → β` to a map `α → β`. -/ +def Ici_extend (f : Ici a → β) : α → β := f ∘ proj_Ici a + +/-- Extend a function `]-∞, b] → β` to a map `α → β`. -/ +def Iic_extend (f : Iic b → β) : α → β := f ∘ proj_Iic b + /-- Extend a function `[a, b] → β` to a map `α → β`. -/ def Icc_extend {a b : α} (h : a ≤ b) (f : Icc a b → β) : α → β := f ∘ proj_Icc a b h +@[simp] lemma Ici_extend_apply (f : Ici a → β) (x : α) : + Ici_extend f x = f ⟨max a x, le_max_left _ _⟩ := rfl +@[simp] lemma Iic_extend_apply (f : Iic b → β) (x : α) : + Iic_extend f x = f ⟨min b x, min_le_left _ _⟩ := rfl +lemma Icc_extend_apply (h : a ≤ b) (f : Icc a b → β) (x : α) : + Icc_extend h f x = f ⟨max a (min b x), le_max_left _ _, max_le h (min_le_left _ _)⟩ := rfl + +@[simp] lemma range_Ici_extend (f : Ici a → β) : range (Ici_extend f) = range f := +by simp only [Ici_extend, range_comp f, range_proj_Ici, range_id'] + +@[simp] lemma range_Iic_extend (f : Iic b → β) : range (Iic_extend f) = range f := +by simp only [Iic_extend, range_comp f, range_proj_Iic, range_id'] + @[simp] lemma Icc_extend_range (f : Icc a b → β) : range (Icc_extend h f) = range f := by simp only [Icc_extend, range_comp f, range_proj_Icc, range_id'] +lemma Ici_extend_of_le (f : Ici a → β) (hx : x ≤ a) : Ici_extend f x = f ⟨a, le_rfl⟩ := +congr_arg f $ proj_Ici_of_le hx + +lemma Iic_extend_of_le (f : Iic b → β) (hx : b ≤ x) : Iic_extend f x = f ⟨b, le_rfl⟩ := +congr_arg f $ proj_Iic_of_le hx + lemma Icc_extend_of_le_left (f : Icc a b → β) (hx : x ≤ a) : Icc_extend h f x = f ⟨a, left_mem_Icc.2 h⟩ := congr_arg f $ proj_Icc_of_le_left h hx -@[simp] lemma Icc_extend_left (f : Icc a b → β) : - Icc_extend h f a = f ⟨a, left_mem_Icc.2 h⟩ := -Icc_extend_of_le_left h f le_rfl - lemma Icc_extend_of_right_le (f : Icc a b → β) (hx : b ≤ x) : Icc_extend h f x = f ⟨b, right_mem_Icc.2 h⟩ := congr_arg f $ proj_Icc_of_right_le h hx +@[simp] lemma Ici_extend_self (f : Ici a → β) : Ici_extend f a = f ⟨a, le_rfl⟩ := +Ici_extend_of_le f le_rfl + +@[simp] lemma Iic_extend_self (f : Iic b → β) : Iic_extend f b = f ⟨b, le_rfl⟩ := +Iic_extend_of_le f le_rfl + +@[simp] lemma Icc_extend_left (f : Icc a b → β) : + Icc_extend h f a = f ⟨a, left_mem_Icc.2 h⟩ := +Icc_extend_of_le_left h f le_rfl + @[simp] lemma Icc_extend_right (f : Icc a b → β) : Icc_extend h f b = f ⟨b, right_mem_Icc.2 h⟩ := Icc_extend_of_right_le h f le_rfl +lemma Ici_extend_of_mem (f : Ici a → β) (hx : x ∈ Ici a) : Ici_extend f x = f ⟨x, hx⟩ := +congr_arg f $ proj_Ici_of_mem hx + +lemma Iic_extend_of_mem (f : Iic b → β) (hx : x ∈ Iic b) : Iic_extend f x = f ⟨x, hx⟩ := +congr_arg f $ proj_Iic_of_mem hx + lemma Icc_extend_of_mem (f : Icc a b → β) (hx : x ∈ Icc a b) : Icc_extend h f x = f ⟨x, hx⟩ := congr_arg f $ proj_Icc_of_mem h hx +@[simp] lemma Ici_extend_coe (f : Ici a → β) (x : Ici a) : Ici_extend f x = f x := +congr_arg f $ proj_Ici_coe x + +@[simp] lemma Iic_extend_coe (f : Iic b → β) (x : Iic b) : Iic_extend f x = f x := +congr_arg f $ proj_Iic_coe x + @[simp] lemma Icc_extend_coe (f : Icc a b → β) (x : Icc a b) : Icc_extend h f x = f x := congr_arg f $ proj_Icc_coe h x @@ -132,11 +208,37 @@ end set open set -variables [preorder β] {a b : α} (h : a ≤ b) {f : Icc a b → β} +variables [preorder β] {s t : set α} {a b : α} (h : a ≤ b) {f : Icc a b → β} + +protected lemma monotone.Ici_extend {f : Ici a → β} (hf : monotone f) : monotone (Ici_extend f) := +hf.comp monotone_proj_Ici + +protected lemma monotone.Iic_extend {f : Iic b → β} (hf : monotone f) : monotone (Iic_extend f) := +hf.comp monotone_proj_Iic -lemma monotone.Icc_extend (hf : monotone f) : monotone (Icc_extend h f) := +protected lemma monotone.Icc_extend (hf : monotone f) : monotone (Icc_extend h f) := hf.comp $ monotone_proj_Icc h +lemma strict_mono.strict_mono_on_Ici_extend {f : Ici a → β} (hf : strict_mono f) : + strict_mono_on (Ici_extend f) (Ici a) := +hf.comp_strict_mono_on strict_mono_on_proj_Ici + +lemma strict_mono.strict_mono_on_Iic_extend {f : Iic b → β} (hf : strict_mono f) : + strict_mono_on (Iic_extend f) (Iic b) := +hf.comp_strict_mono_on strict_mono_on_proj_Iic + lemma strict_mono.strict_mono_on_Icc_extend (hf : strict_mono f) : strict_mono_on (Icc_extend h f) (Icc a b) := hf.comp_strict_mono_on (strict_mono_on_proj_Icc h) + +protected lemma set.ord_connected.Ici_extend {s : set (Ici a)} (hs : s.ord_connected) : + {x | Ici_extend (∈ s) x}.ord_connected := +⟨λ x hx y hy z hz, hs.out hx hy ⟨max_le_max le_rfl hz.1, max_le_max le_rfl hz.2⟩⟩ + +protected lemma set.ord_connected.Iic_extend {s : set (Iic b)} (hs : s.ord_connected) : + {x | Iic_extend (∈ s) x}.ord_connected := +⟨λ x hx y hy z hz, hs.out hx hy ⟨min_le_min le_rfl hz.1, min_le_min le_rfl hz.2⟩⟩ + +protected lemma set.ord_connected.restrict (hs : s.ord_connected) : + {x | restrict t (∈ s) x}.ord_connected := +⟨λ x hx y hy z hz, hs.out hx hy hz⟩ From 3a2b5524a138b5d0b818b858b516d4ac8a484b03 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 12 Jul 2023 12:15:34 +0000 Subject: [PATCH 1241/1291] feat(data/fin/basic): extra instances that cover `fin 0` (#18970) These apply to `fin 0`, unlike the `comm_ring` instance which needs `ne_zero n`. Co-authored-by: Yury G. Kudryashov --- src/data/fin/basic.lean | 26 ++++++++++++++++++++++++-- src/data/zmod/defs.lean | 18 +++++++++++++++--- 2 files changed, 39 insertions(+), 5 deletions(-) diff --git a/src/data/fin/basic.lean b/src/data/fin/basic.lean index a5ba4089bbf05..1579e1a3837ba 100644 --- a/src/data/fin/basic.lean +++ b/src/data/fin/basic.lean @@ -423,6 +423,11 @@ by rcases n with _|_|n; simp [is_empty.subsingleton, unique.subsingleton, not_su section monoid +instance add_comm_semigroup (n : ℕ) : add_comm_semigroup (fin n) := +{ add := (+), + add_assoc := by simp [eq_iff_veq, add_def, add_assoc], + add_comm := by simp [eq_iff_veq, add_def, add_comm] } + @[simp] protected lemma add_zero [ne_zero n] (k : fin n) : k + 0 = k := by simp [eq_iff_veq, add_def, mod_eq_of_lt (is_lt k)] @@ -431,11 +436,10 @@ by simp [eq_iff_veq, add_def, mod_eq_of_lt (is_lt k)] instance add_comm_monoid (n : ℕ) [ne_zero n] : add_comm_monoid (fin n) := { add := (+), - add_assoc := by simp [eq_iff_veq, add_def, add_assoc], zero := 0, zero_add := fin.zero_add, add_zero := fin.add_zero, - add_comm := by simp [eq_iff_veq, add_def, add_comm] } + ..fin.add_comm_semigroup n } instance [ne_zero n] : add_monoid_with_one (fin n) := { one := 1, @@ -1363,6 +1367,24 @@ instance (n : ℕ) [ne_zero n] : add_comm_group (fin n) := ..fin.add_comm_monoid n, ..fin.has_neg n } +/-- Note this is more general than `fin.add_comm_group` as it applies (vacuously) to `fin 0` too. -/ +instance (n : ℕ) : has_involutive_neg (fin n) := +{ neg := has_neg.neg, + neg_neg := nat.cases_on n fin_zero_elim (λ i, neg_neg) } + +/-- Note this is more general than `fin.add_comm_group` as it applies (vacuously) to `fin 0` too. -/ +instance (n : ℕ) : is_cancel_add (fin n) := +{ add_left_cancel := nat.cases_on n fin_zero_elim (λ i _ _ _, add_left_cancel), + add_right_cancel := nat.cases_on n fin_zero_elim (λ i _ _ _, add_right_cancel) } + +/-- Note this is more general than `fin.add_comm_group` as it applies (vacuously) to `fin 0` too. -/ +instance (n : ℕ) : add_left_cancel_semigroup (fin n) := +{ ..fin.add_comm_semigroup n, .. fin.is_cancel_add n } + +/-- Note this is more general than `fin.add_comm_group` as it applies (vacuously) to `fin 0` too. -/ +instance (n : ℕ) : add_right_cancel_semigroup (fin n) := +{ ..fin.add_comm_semigroup n, .. fin.is_cancel_add n } + protected lemma coe_neg (a : fin n) : ((-a : fin n) : ℕ) = (n - a) % n := rfl protected lemma coe_sub (a b : fin n) : ((a - b : fin n) : ℕ) = (a + (n - b)) % n := diff --git a/src/data/zmod/defs.lean b/src/data/zmod/defs.lean index a8a6a3732836c..918a0caf9ab08 100644 --- a/src/data/zmod/defs.lean +++ b/src/data/zmod/defs.lean @@ -57,15 +57,27 @@ private lemma left_distrib_aux (n : ℕ) : ∀ a b c : fin n, a * (b + c) = a * ... ≡ (a * b) % n + (a * c) % n [MOD n] : (nat.mod_modeq _ _).symm.add (nat.mod_modeq _ _).symm) +instance (n : ℕ) : distrib (fin n) := +{ left_distrib := left_distrib_aux n, + right_distrib := λ a b c, by rw [mul_comm, left_distrib_aux, mul_comm _ b, mul_comm]; refl, + ..fin.add_comm_semigroup n, + ..fin.comm_semigroup n } + /-- Commutative ring structure on `fin n`. -/ instance (n : ℕ) [ne_zero n] : comm_ring (fin n) := { one_mul := fin.one_mul, mul_one := fin.mul_one, - left_distrib := left_distrib_aux n, - right_distrib := λ a b c, by rw [mul_comm, left_distrib_aux, mul_comm _ b, mul_comm]; refl, ..fin.add_monoid_with_one, ..fin.add_comm_group n, - ..fin.comm_semigroup n } + ..fin.comm_semigroup n, + ..fin.distrib n } + +/-- Note this is more general than `fin.comm_ring` as it applies (vacuously) to `fin 0` too. -/ +instance (n : ℕ) : has_distrib_neg (fin n) := +{ neg := has_neg.neg, + mul_neg := nat.cases_on n fin_zero_elim $ λ i, mul_neg, + neg_mul := nat.cases_on n fin_zero_elim $ λ i, neg_mul, + ..fin.has_involutive_neg n } end fin From 92bd7b1ffeb306a89f450bee126ddd8a284c259d Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 12 Jul 2023 15:28:30 +0000 Subject: [PATCH 1242/1291] feat(analysis/convex): convexity of n-ary sums (#18943) --- src/analysis/convex/basic.lean | 30 +++++++++++++++++++ src/analysis/convex/combination.lean | 20 +++++++++++++ src/analysis/convex/hull.lean | 3 ++ src/measure_theory/measure/haar/of_basis.lean | 15 ++-------- 4 files changed, 55 insertions(+), 13 deletions(-) diff --git a/src/analysis/convex/basic.lean b/src/analysis/convex/basic.lean index 6825f10078981..d1520b6a6ab74 100644 --- a/src/analysis/convex/basic.lean +++ b/src/analysis/convex/basic.lean @@ -152,6 +152,9 @@ convex_iff_pairwise_pos.mpr (h.pairwise _) lemma convex_singleton (c : E) : convex 𝕜 ({c} : set E) := subsingleton_singleton.convex +lemma convex_zero : convex 𝕜 (0 : set E) := +convex_singleton _ + lemma convex_segment (x y : E) : convex 𝕜 [x -[𝕜] y] := begin rintro p ⟨ap, bp, hap, hbp, habp, rfl⟩ q ⟨aq, bq, haq, hbq, habq, rfl⟩ a b ha hb hab, @@ -190,6 +193,33 @@ hs.linear_preimage $ hf.mk' f lemma convex.add {t : set E} (hs : convex 𝕜 s) (ht : convex 𝕜 t) : convex 𝕜 (s + t) := by { rw ← add_image_prod, exact (hs.prod ht).is_linear_image is_linear_map.is_linear_map_add } +variables (𝕜 E) + +/-- The convex sets form an additive submonoid under pointwise addition. -/ +def convex_add_submonoid : add_submonoid (set E) := +{ carrier := {s : set E | convex 𝕜 s}, + zero_mem' := convex_zero, + add_mem' := λ s t, convex.add } + +@[simp, norm_cast] +lemma coe_convex_add_submonoid : ↑(convex_add_submonoid 𝕜 E) = {s : set E | convex 𝕜 s} := rfl + +variables {𝕜 E} + +@[simp] lemma mem_convex_add_submonoid {s : set E} : + s ∈ convex_add_submonoid 𝕜 E ↔ convex 𝕜 s := +iff.rfl + +lemma convex_list_sum {l : list (set E)} (h : ∀ i ∈ l, convex 𝕜 i) : convex 𝕜 l.sum := +(convex_add_submonoid 𝕜 E).list_sum_mem h + +lemma convex_multiset_sum {s : multiset (set E)} (h : ∀ i ∈ s, convex 𝕜 i) : convex 𝕜 s.sum := +(convex_add_submonoid 𝕜 E).multiset_sum_mem _ h + +lemma convex_sum {ι} {s : finset ι} (t : ι → set E) (h : ∀ i ∈ s, convex 𝕜 (t i)) : + convex 𝕜 (∑ i in s, t i) := +(convex_add_submonoid 𝕜 E).sum_mem h + lemma convex.vadd (hs : convex 𝕜 s) (z : E) : convex 𝕜 (z +ᵥ s) := by { simp_rw [←image_vadd, vadd_eq_add, ←singleton_add], exact (convex_singleton _).add hs } diff --git a/src/analysis/convex/combination.lean b/src/analysis/convex/combination.lean index 8cc8ae375c102..15443b1b1fa4f 100644 --- a/src/analysis/convex/combination.lean +++ b/src/analysis/convex/combination.lean @@ -418,9 +418,29 @@ lemma convex_hull_add (s t : set E) : convex_hull R (s + t) = convex_hull R s + by simp_rw [←image2_add, ←image_prod, is_linear_map.is_linear_map_add.convex_hull_image, convex_hull_prod] +variables (R E) +/-- `convex_hull` is an additive monoid morphism under pointwise addition. -/ +@[simps] +def convex_hull_add_monoid_hom : set E →+ set E := +{ to_fun := convex_hull R, + map_add' := convex_hull_add, + map_zero' := convex_hull_zero } +variables {R E} + lemma convex_hull_sub (s t : set E) : convex_hull R (s - t) = convex_hull R s - convex_hull R t := by simp_rw [sub_eq_add_neg, convex_hull_add, convex_hull_neg] +lemma convex_hull_list_sum (l : list (set E)) : convex_hull R l.sum = (l.map $ convex_hull R).sum := +map_list_sum (convex_hull_add_monoid_hom R E) l + +lemma convex_hull_multiset_sum (s : multiset (set E)) : + convex_hull R s.sum = (s.map $ convex_hull R).sum := +map_multiset_sum (convex_hull_add_monoid_hom R E) s + +lemma convex_hull_sum {ι} (s : finset ι) (t : ι → set E) : + convex_hull R (∑ i in s, t i) = ∑ i in s, convex_hull R (t i):= +map_sum (convex_hull_add_monoid_hom R E) _ _ + /-! ### `std_simplex` -/ variables (ι) [fintype ι] {f : ι → R} diff --git a/src/analysis/convex/hull.lean b/src/analysis/convex/hull.lean index efbc727029664..535a17885c56f 100644 --- a/src/analysis/convex/hull.lean +++ b/src/analysis/convex/hull.lean @@ -101,6 +101,9 @@ lemma segment_subset_convex_hull (hx : x ∈ s) (hy : y ∈ s) : segment 𝕜 x @[simp] lemma convex_hull_singleton (x : E) : convex_hull 𝕜 ({x} : set E) = {x} := (convex_singleton x).convex_hull_eq +@[simp] lemma convex_hull_zero : convex_hull 𝕜 (0 : set E) = 0 := +convex_hull_singleton 0 + @[simp] lemma convex_hull_pair (x y : E) : convex_hull 𝕜 {x, y} = segment 𝕜 x y := begin refine (convex_hull_min _ $ convex_segment _ _).antisymm diff --git a/src/measure_theory/measure/haar/of_basis.lean b/src/measure_theory/measure/haar/of_basis.lean index d8aed76d5d53d..dd1d618b23b18 100644 --- a/src/measure_theory/measure/haar/of_basis.lean +++ b/src/measure_theory/measure/haar/of_basis.lean @@ -126,24 +126,13 @@ end lemma convex_parallelepiped (v : ι → E) : convex ℝ (parallelepiped v) := begin rw parallelepiped_eq_sum_segment, - -- TODO: add `convex.sum` to match `convex.add` - let : add_submonoid (set E) := - { carrier := { s | convex ℝ s}, zero_mem' := convex_singleton _, add_mem' := λ x y, convex.add }, - exact this.sum_mem (λ i hi, convex_segment _ _), + exact convex_sum _ (λ i hi, convex_segment _ _), end /-- A `parallelepiped` is the convex hull of its vertices -/ lemma parallelepiped_eq_convex_hull (v : ι → E) : parallelepiped v = convex_hull ℝ (∑ i, {(0 : E), v i}) := -begin - -- TODO: add `convex_hull_sum` to match `convex_hull_add` - let : set E →+ set E := - { to_fun := convex_hull ℝ, - map_zero' := convex_hull_singleton _, - map_add' := convex_hull_add }, - simp_rw [parallelepiped_eq_sum_segment, ←convex_hull_pair], - exact (this.map_sum _ _).symm, -end +by simp_rw [convex_hull_sum, convex_hull_pair, parallelepiped_eq_sum_segment] /-- The axis aligned parallelepiped over `ι → ℝ` is a cuboid. -/ lemma parallelepiped_single [decidable_eq ι] (a : ι → ℝ) : From d9e96a3e3e0894e93e10aff5244f4c96655bac1c Mon Sep 17 00:00:00 2001 From: Devon Tuma Date: Wed, 12 Jul 2023 15:28:31 +0000 Subject: [PATCH 1243/1291] feat(data/list/dedup): Lemmas about `list.dedup` (#19142) Basic lemmas about `dedup` applied with `cons`, `nil`, `head`, and `tail`. Co-authored-by: Devon Tuma --- src/data/finset/basic.lean | 6 ++++++ src/data/list/dedup.lean | 37 +++++++++++++++++++++++++++++++++++++ 2 files changed, 43 insertions(+) diff --git a/src/data/finset/basic.lean b/src/data/finset/basic.lean index 43bca0b7aca26..ae0db3da0b9b8 100644 --- a/src/data/finset/basic.lean +++ b/src/data/finset/basic.lean @@ -2162,6 +2162,9 @@ by ext; simp @[simp] theorem to_finset_eq_empty {m : multiset α} : m.to_finset = ∅ ↔ m = 0 := finset.val_inj.symm.trans multiset.dedup_eq_zero +@[simp] lemma to_finset_nonempty : s.to_finset.nonempty ↔ s ≠ 0 := +by simp only [to_finset_eq_empty, ne.def, finset.nonempty_iff_ne_empty] + @[simp] lemma to_finset_subset : s.to_finset ⊆ t.to_finset ↔ s ⊆ t := by simp only [finset.subset_iff, multiset.subset_iff, multiset.mem_to_finset] @@ -2253,6 +2256,9 @@ by { ext, simp } @[simp] lemma to_finset_eq_empty_iff (l : list α) : l.to_finset = ∅ ↔ l = nil := by cases l; simp +@[simp] lemma to_finset_nonempty_iff (l : list α) : l.to_finset.nonempty ↔ l ≠ [] := +by simp [finset.nonempty_iff_ne_empty] + end list namespace finset diff --git a/src/data/list/dedup.lean b/src/data/list/dedup.lean index 2bfdd74a3ad16..4b42facffb4c6 100644 --- a/src/data/list/dedup.lean +++ b/src/data/list/dedup.lean @@ -57,8 +57,45 @@ theorem subset_dedup (l : list α) : l ⊆ dedup l := theorem nodup_dedup : ∀ l : list α, nodup (dedup l) := pairwise_pw_filter +theorem head_dedup [inhabited α] (l : list α) : + l.dedup.head = if l.head ∈ l.tail then l.tail.dedup.head else l.head := +match l with +| [] := rfl +| (a :: l) := by { by_cases ha : a ∈ l; simp [ha, list.dedup_cons_of_mem] } +end + +theorem tail_dedup [inhabited α] (l : list α) : + l.dedup.tail = if l.head ∈ l.tail then l.tail.dedup.tail else l.tail.dedup := +match l with +| [] := rfl +| (a :: l) := by { by_cases ha : a ∈ l; simp [ha, list.dedup_cons_of_mem] } +end + theorem dedup_eq_self {l : list α} : dedup l = l ↔ nodup l := pw_filter_eq_self +theorem dedup_eq_cons (l : list α) (a : α) (l' : list α) : + l.dedup = a :: l' ↔ a ∈ l ∧ a ∉ l' ∧ l.dedup.tail = l' := +begin + refine ⟨λ h, _, λ h, _⟩, + { refine ⟨mem_dedup.1 (h.symm ▸ mem_cons_self _ _), λ ha, _, by rw [h, tail_cons]⟩, + have : count a l.dedup ≤ 1 := nodup_iff_count_le_one.1 (nodup_dedup l) a, + rw [h, count_cons_self, add_le_iff_nonpos_left] at this, + exact (not_le_of_lt (count_pos.2 ha) this) }, + { have := @cons_head_tail α ⟨a⟩ _ (ne_nil_of_mem (mem_dedup.2 h.1)), + have hal : a ∈ l.dedup := mem_dedup.2 h.1, + rw [← this, mem_cons_iff, or_iff_not_imp_right] at hal, + exact this ▸ h.2.2.symm ▸ (cons_eq_cons.2 ⟨(hal (h.2.2.symm ▸ h.2.1)).symm, rfl⟩) } +end + +@[simp] theorem dedup_eq_nil (l : list α) : l.dedup = [] ↔ l = [] := +begin + induction l with a l hl, + { exact iff.rfl }, + { by_cases h : a ∈ l, + { simp only [list.dedup_cons_of_mem h, hl, list.ne_nil_of_mem h] }, + { simp only [list.dedup_cons_of_not_mem h, list.cons_ne_nil] } } +end + protected lemma nodup.dedup {l : list α} (h : l.nodup) : l.dedup = l := list.dedup_eq_self.2 h From d3b54a9f9613af6ae818c7fee31872db11c4d928 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 12 Jul 2023 15:28:32 +0000 Subject: [PATCH 1244/1291] fix(tactic/interval_cases): instantiate metavars (#19232) The test included in this commit fails without this change. This does not need forward-porting, the test already passes in Lean 4. --- src/tactic/interval_cases.lean | 12 ++++++------ test/interval_cases.lean | 7 +++++++ 2 files changed, 13 insertions(+), 6 deletions(-) diff --git a/src/tactic/interval_cases.lean b/src/tactic/interval_cases.lean index 9cffa15e8d41a..5fd7bfac4410f 100644 --- a/src/tactic/interval_cases.lean +++ b/src/tactic/interval_cases.lean @@ -47,14 +47,14 @@ return that proof. -- We use `expr.to_rat` merely to decide if an `expr` is an explicit number. -- It would be more natural to use `expr.to_int`, but that hasn't been implemented. meta def gives_upper_bound (n e : expr) : tactic expr := -do t ← infer_type e, +do t ← infer_type e >>= instantiate_mvars, match t with | `(%%n' < %%b) := do guard (n = n'), b ← b.to_rat, return e | `(%%b > %%n') := do guard (n = n'), b ← b.to_rat, return e | `(%%n' ≤ %%b) := do guard (n = n'), b ← b.to_rat, - tn ← infer_type n, + tn ← infer_type n >>= instantiate_mvars, match tn with | `(ℕ) := to_expr ``(nat.lt_add_one_iff.mpr %%e) | `(ℕ+) := to_expr ``(pnat.lt_add_one_iff.mpr %%e) @@ -64,7 +64,7 @@ do t ← infer_type e, | `(%%b ≥ %%n') := do guard (n = n'), b ← b.to_rat, - tn ← infer_type n, + tn ← infer_type n >>= instantiate_mvars, match tn with | `(ℕ) := to_expr ``(nat.lt_add_one_iff.mpr %%e) | `(ℕ+) := to_expr ``(pnat.lt_add_one_iff.mpr %%e) @@ -80,14 +80,14 @@ for some explicit `b`, return that proof. -/ meta def gives_lower_bound (n e : expr) : tactic expr := -do t ← infer_type e, +do t ← infer_type e >>= instantiate_mvars, match t with | `(%%n' ≥ %%b) := do guard (n = n'), b ← b.to_rat, return e | `(%%b ≤ %%n') := do guard (n = n'), b ← b.to_rat, return e | `(%%n' > %%b) := do guard (n = n'), b ← b.to_rat, - tn ← infer_type n, + tn ← infer_type n >>= instantiate_mvars, match tn with | `(ℕ) := to_expr ``(nat.add_one_le_iff.mpr %%e) | `(ℕ+) := to_expr ``(pnat.add_one_le_iff.mpr %%e) @@ -97,7 +97,7 @@ do t ← infer_type e, | `(%%b < %%n') := do guard (n = n'), b ← b.to_rat, - tn ← infer_type n, + tn ← infer_type n >>= instantiate_mvars, match tn with | `(ℕ) := to_expr ``(nat.add_one_le_iff.mpr %%e) | `(ℕ+) := to_expr ``(pnat.add_one_le_iff.mpr %%e) diff --git a/test/interval_cases.lean b/test/interval_cases.lean index 7e9c8e528b0a1..5c9f6074378f2 100644 --- a/test/interval_cases.lean +++ b/test/interval_cases.lean @@ -136,6 +136,13 @@ begin exact h, end +example : ∀ y, y ≤ 3 → true := +begin + refine λ y hy, _, + interval_cases y, + all_goals { trivial }, +end + /- Sadly, this one doesn't work, reporting: `deep recursion was detected at 'expression equality test'` From 41bef4ae1254365bc190aee63b947674d2977f01 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 12 Jul 2023 18:27:42 +0000 Subject: [PATCH 1245/1291] feat(analysis/normed/group/basic): norm lemmas for lipschitz and antilipschitz (#19103) This also corrects some nonsense names produced by to_additive. --- src/analysis/calculus/fderiv/basic.lean | 2 +- src/analysis/normed/group/basic.lean | 30 ++++++++++++++++--- .../normed_space/continuous_linear_map.lean | 2 +- 3 files changed, 28 insertions(+), 6 deletions(-) diff --git a/src/analysis/calculus/fderiv/basic.lean b/src/analysis/calculus/fderiv/basic.lean index b7fcebc366f5d..ca62479ba25cc 100644 --- a/src/analysis/calculus/fderiv/basic.lean +++ b/src/analysis/calculus/fderiv/basic.lean @@ -693,7 +693,7 @@ lemma has_fderiv_at_filter.is_O_sub_rev (hf : has_fderiv_at_filter f f' x L) {C} (λ x', x' - x) =O[L] (λ x', f x' - f x) := have (λ x', x' - x) =O[L] (λ x', f' (x' - x)), from is_O_iff.2 ⟨C, eventually_of_forall $ λ x', - add_monoid_hom_class.bound_of_antilipschitz f' hf' _⟩, + zero_hom_class.bound_of_antilipschitz f' hf' _⟩, (this.trans (hf.trans_is_O this).right_is_O_add).congr (λ _, rfl) (λ _, sub_add_cancel _ _) end continuous diff --git a/src/analysis/normed/group/basic.lean b/src/analysis/normed/group/basic.lean index 7b25f1a4fc002..6fe45df80f9c3 100644 --- a/src/analysis/normed/group/basic.lean +++ b/src/analysis/normed/group/basic.lean @@ -387,7 +387,8 @@ by simpa [dist_eq_norm_div] using dist_triangle a 1 b ‖a₁ / a₂‖ ≤ r₁ + r₂ := (norm_div_le a₁ a₂).trans $ add_le_add H₁ H₂ -@[to_additive] lemma dist_le_norm_mul_norm (a b : E) : dist a b ≤ ‖a‖ + ‖b‖ := +@[to_additive dist_le_norm_add_norm] lemma dist_le_norm_add_norm' (a b : E) : + dist a b ≤ ‖a‖ + ‖b‖ := by { rw dist_eq_norm_div, apply norm_div_le } @[to_additive abs_norm_sub_norm_le] lemma abs_norm_sub_norm_le' (a b : E) : |‖a‖ - ‖b‖| ≤ ‖a / b‖ := @@ -656,9 +657,29 @@ by rw [emetric.mem_ball, edist_eq_coe_nnnorm'] antilipschitz_with K f := antilipschitz_with.of_le_mul_dist $ λ x y, by simpa only [dist_eq_norm_div, map_div] using h (x / y) -@[to_additive] lemma monoid_hom_class.bound_of_antilipschitz [monoid_hom_class 𝓕 E F] (f : 𝓕) +@[to_additive lipschitz_with.norm_le_mul] +lemma lipschitz_with.norm_le_mul' {f : E → F} + {K : ℝ≥0} (h : lipschitz_with K f) (hf : f 1 = 1) (x) : ‖f x‖ ≤ K * ‖x‖ := +by simpa only [dist_one_right, hf] using h.dist_le_mul x 1 + +@[to_additive lipschitz_with.nnorm_le_mul] +lemma lipschitz_with.nnorm_le_mul' {f : E → F} + {K : ℝ≥0} (h : lipschitz_with K f) (hf : f 1 = 1) (x) : ‖f x‖₊ ≤ K * ‖x‖₊ := +h.norm_le_mul' hf x + +@[to_additive antilipschitz_with.le_mul_norm] +lemma antilipschitz_with.le_mul_norm' {f : E → F} + {K : ℝ≥0} (h : antilipschitz_with K f) (hf : f 1 = 1) (x) : ‖x‖ ≤ K * ‖f x‖ := +by simpa only [dist_one_right, hf] using h.le_mul_dist x 1 + +@[to_additive antilipschitz_with.le_mul_nnnorm] +lemma antilipschitz_with.le_mul_nnnorm' {f : E → F} + {K : ℝ≥0} (h : antilipschitz_with K f) (hf : f 1 = 1) (x) : ‖x‖₊ ≤ K * ‖f x‖₊ := +h.le_mul_norm' hf x + +@[to_additive] lemma one_hom_class.bound_of_antilipschitz [one_hom_class 𝓕 E F] (f : 𝓕) {K : ℝ≥0} (h : antilipschitz_with K f) (x) : ‖x‖ ≤ K * ‖f x‖ := -by simpa only [dist_one_right, map_one] using h.le_mul_dist x 1 +h.le_mul_nnnorm' (map_one f) x end nnnorm @@ -1285,7 +1306,8 @@ end (hg : lipschitz_with Kg (g / f)) (hK : Kg < Kf⁻¹) : antilipschitz_with (Kf⁻¹ - Kg)⁻¹ g := by simpa only [pi.div_apply, mul_div_cancel'_right] using hf.mul_lipschitz_with hg hK -@[to_additive] lemma le_mul_norm_div {f : E → F} (hf : antilipschitz_with K f) (x y : E) : +@[to_additive le_mul_norm_sub] +lemma le_mul_norm_div {f : E → F} (hf : antilipschitz_with K f) (x y : E) : ‖x / y‖ ≤ K * ‖f x / f y‖ := by simp [← dist_eq_norm_div, hf.le_mul_dist x y] diff --git a/src/analysis/normed_space/continuous_linear_map.lean b/src/analysis/normed_space/continuous_linear_map.lean index 87ad5e2c41002..a46af72230230 100644 --- a/src/analysis/normed_space/continuous_linear_map.lean +++ b/src/analysis/normed_space/continuous_linear_map.lean @@ -105,7 +105,7 @@ add_monoid_hom_class.antilipschitz_of_bound _ h lemma bound_of_antilipschitz (f : E →SL[σ] F) {K : ℝ≥0} (h : antilipschitz_with K f) (x) : ‖x‖ ≤ K * ‖f x‖ := -add_monoid_hom_class.bound_of_antilipschitz _ h x +zero_hom_class.bound_of_antilipschitz _ h x end continuous_linear_map From b01d6eb9d0a308807af54319b264d0994b91774b Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Pol=27tta=20/=20Miyahara=20K=C5=8D?= Date: Wed, 12 Jul 2023 18:27:44 +0000 Subject: [PATCH 1246/1291] fix(control/traversable/derive): the `functor` deriving handler makes weird definitions for inductive types which have recursive arguments separated by a non-recursive argument (#19228) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit I found this bug during the porting. The current deriving handler doesn't work for this. ```lean @[derive [traversable,is_lawful_traversable]] inductive my_tree' (α : Type) | leaf : my_tree' | node : my_tree' → α → my_tree' → my_tree' ``` This is because the `functor` deriving handler makes weird definitions for inductive types which have recursive arguments separated by a non-recursive argument. Fortunatelly, the cause of this bug is just a mistake of the argument in `control.traversable.derive`. --- src/control/traversable/derive.lean | 2 +- test/traversable.lean | 5 +++++ 2 files changed, 6 insertions(+), 1 deletion(-) diff --git a/src/control/traversable/derive.lean b/src/control/traversable/derive.lean index 3bd145a69f05f..3545bc23feb9c 100644 --- a/src/control/traversable/derive.lean +++ b/src/control/traversable/derive.lean @@ -50,7 +50,7 @@ meta def map_constructor (c n : name) (f α β : expr) do g ← target, (_, args') ← mmap_accuml (λ (x : list expr) (y : bool × expr), if y.1 then pure (x.tail,x.head) - else prod.mk rec_call <$> map_field n g.app_fn f α β y.2) rec_call args₁, + else prod.mk x <$> map_field n g.app_fn f α β y.2) rec_call args₁, constr ← mk_const c, let r := constr.mk_app (args₀ ++ args'), return r diff --git a/test/traversable.lean b/test/traversable.lean index 6adc03f87b411..9751f19392291 100644 --- a/test/traversable.lean +++ b/test/traversable.lean @@ -46,6 +46,11 @@ inductive my_tree (α : Type) | leaf : my_tree | node : my_tree → my_tree → α → my_tree +@[derive [traversable,is_lawful_traversable]] +inductive my_tree' (α : Type) +| leaf : my_tree' +| node : my_tree' → α → my_tree' → my_tree' + section open my_tree (hiding traverse) From baa88307f3e699fa7054ef04ec79fa4f056169cb Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Wed, 12 Jul 2023 21:38:36 +0000 Subject: [PATCH 1247/1291] feat(analysis/inner_product_space/of_norm): Create an inner product from a norm (#4798) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit A normed space respecting the polarization identity is an inner product space. Co-authored-by: Frédéric Dupuis Co-authored-by: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Co-authored-by: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Co-authored-by: Yaël Dillies --- docs/references.bib | 14 + src/analysis/inner_product_space/of_norm.lean | 379 ++++++++++++++++++ src/data/is_R_or_C/basic.lean | 6 + 3 files changed, 399 insertions(+) create mode 100644 src/analysis/inner_product_space/of_norm.lean diff --git a/docs/references.bib b/docs/references.bib index d97389ddd156d..03a341208d7bb 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -1406,6 +1406,20 @@ @Book{ james1999 url = {https://doi.org/10.1007/978-1-4471-3994-2} } +@Article{ Jordan1935, + title = "On inner products in linear, metric spaces", + author = "Jordan, P. and von Neumann, J.", + fjournal = {{Annals of Mathematics}}, + journal = "Ann. Math.", + volume = 36, + number = 3, + pages = "719-723", + month = jul, + year = 1935, + url = "http://www.mathematik.uni-muenchen.de/~michel/jordan-von_neumann_-_parallelogram_identity.pdf", + doi = {10.2307/1968653} +} + @Article{ joyal1977, author = {André Joyal}, title = {Remarques sur la théorie des jeux à deux personnes}, diff --git a/src/analysis/inner_product_space/of_norm.lean b/src/analysis/inner_product_space/of_norm.lean new file mode 100644 index 0000000000000..be0bf32167923 --- /dev/null +++ b/src/analysis/inner_product_space/of_norm.lean @@ -0,0 +1,379 @@ +/- +Copyright (c) 2020 Heather Macbeth. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Heather Macbeth +-/ + +import topology.algebra.algebra +import analysis.inner_product_space.basic + +/-! +# Inner product space derived from a norm + +This file defines an `inner_product_space` instance from a norm that respects the +parallellogram identity. The parallelogram identity is a way to express the inner product of `x` and +`y` in terms of the norms of `x`, `y`, `x + y`, `x - y`. + +## Main results + +- `inner_product_space.of_norm`: a normed space whose norm respects the parallellogram identity, + can be seen as an inner product space. + +## Implementation notes + +We define `inner_` + +$$\langle x, y \rangle := \frac{1}{4} (‖x + y‖^2 - ‖x - y‖^2 + i ‖ix + y‖ ^ 2 - i ‖ix - y‖^2)$$ + +and use the parallelogram identity + +$$‖x + y‖^2 + ‖x - y‖^2 = 2 (‖x‖^2 + ‖y‖^2)$$ + +to prove it is an inner product, i.e., that it is conjugate-symmetric (`inner_.conj_symm`) and +linear in the first argument. `add_left` is proved by judicious application of the parallelogram +identity followed by tedious arithmetic. `smul_left` is proved step by step, first noting that +$\langle λ x, y \rangle = λ \langle x, y \rangle$ for $λ ∈ ℕ$, $λ = -1$, hence $λ ∈ ℤ$ and $λ ∈ ℚ$ +by arithmetic. Then by continuity and the fact that ℚ is dense in ℝ, the same is true for ℝ. +The case of ℂ then follows by applying the result for ℝ and more arithmetic. + +## TODO + +Move upstream to `analysis.inner_product_space.basic`. + +## References + +- [Jordan, P. and von Neumann, J., *On inner products in linear, metric spaces*][Jordan1935] +- https://math.stackexchange.com/questions/21792/norms-induced-by-inner-products-and-the-parallelogram-law +- https://math.dartmouth.edu/archive/m113w10/public_html/jordan-vneumann-thm.pdf + +## Tags + +inner product space, Hilbert space, norm +-/ + +open is_R_or_C +open_locale complex_conjugate + +variables {𝕜 : Type*} [is_R_or_C 𝕜] (E : Type*) [normed_add_comm_group E] + +/-- Predicate for the parallelogram identity to hold in a normed group. This is a scalar-less +version of `inner_product_space`. If you have an `inner_product_spaceable` assumption, you can +locally upgrade that to `inner_product_space 𝕜 E` using `casesI nonempty_inner_product_space 𝕜 E`. +-/ +class inner_product_spaceable : Prop := +(parallelogram_identity : + ∀ x y : E, ‖x + y‖ * ‖x + y‖ + ‖x - y‖ * ‖x - y‖ = 2 * (‖x‖ * ‖x‖ + ‖y‖ * ‖y‖)) + +variables (𝕜) {E} + +lemma inner_product_space.to_inner_product_spaceable [inner_product_space 𝕜 E] : + inner_product_spaceable E := +⟨parallelogram_law_with_norm 𝕜⟩ + +@[priority 100] -- See note [lower instance priority] +instance inner_product_space.to_inner_product_spaceable_of_real [inner_product_space ℝ E] : + inner_product_spaceable E := +⟨parallelogram_law_with_norm ℝ⟩ + +variables [normed_space 𝕜 E] + +local notation `𝓚` := algebra_map ℝ 𝕜 + +/-- Auxiliary definition of the inner product derived from the norm. -/ +private noncomputable def inner_ (x y : E) : 𝕜 := +4⁻¹ * ((𝓚 ‖x + y‖) * (𝓚 ‖x + y‖) - (𝓚 ‖x - y‖) * (𝓚 ‖x - y‖) + + (I:𝕜) * (𝓚 ‖(I:𝕜) • x + y‖) * (𝓚 ‖(I:𝕜) • x + y‖) + - (I:𝕜) * (𝓚 ‖(I:𝕜) • x - y‖) * (𝓚 ‖(I:𝕜) • x - y‖)) + +namespace inner_product_spaceable + +variables {𝕜} (E) + +/-- Auxiliary definition for the `add_left` property -/ +private def inner_prop (r : 𝕜) : Prop := ∀ x y : E, inner_ 𝕜 (r • x) y = conj r * inner_ 𝕜 x y + +variables {E} + +lemma inner_prop_neg_one : inner_prop E ((-1 : ℤ) : 𝕜) := +begin + intros x y, + simp only [inner_, neg_mul_eq_neg_mul, one_mul, int.cast_one, one_smul, ring_hom.map_one, + map_neg, int.cast_neg, neg_smul, neg_one_mul], + rw neg_mul_comm, + congr' 1, + have h₁ : ‖-x - y‖ = ‖x + y‖, + { rw [←neg_add', norm_neg], }, + have h₂ : ‖-x + y‖ = ‖x - y‖, + { rw [←neg_sub, norm_neg, sub_eq_neg_add], }, + have h₃ : ‖(I : 𝕜) • (-x) + y‖ = ‖(I : 𝕜) • x - y‖, + { rw [←neg_sub, norm_neg, sub_eq_neg_add, ←smul_neg], }, + have h₄ : ‖(I : 𝕜) • (-x) - y‖ = ‖(I : 𝕜) • x + y‖, + { rw [smul_neg, ←neg_add', norm_neg] }, + rw [h₁, h₂, h₃, h₄], + ring, +end + +lemma continuous.inner_ {f g : ℝ → E} (hf : continuous f) (hg : continuous g) : + continuous (λ x, inner_ 𝕜 (f x) (g x)) := +by { unfold inner_, continuity } + +lemma inner_.norm_sq (x : E) : ‖x‖ ^ 2 = re (inner_ 𝕜 x x) := +begin + simp only [inner_], + have h₁ : norm_sq (4 : 𝕜) = 16, + { have : ((4 : ℝ) : 𝕜) = (4 : 𝕜), + { simp only [of_real_one, of_real_bit0] }, + rw [←this, norm_sq_eq_def', + is_R_or_C.norm_of_nonneg (by norm_num : (0 : ℝ) ≤ 4)], + norm_num }, + have h₂ : ‖x + x‖ = 2 * ‖x‖, + { rw [←two_smul 𝕜, norm_smul, is_R_or_C.norm_two] }, + simp only [inner, h₁, h₂, one_im, bit0_zero, add_zero, norm_zero, I_re, of_real_im, map_add, + bit0_im, zero_div, zero_mul, add_monoid_hom.map_neg, of_real_re, map_sub, sub_zero, inv_re, + one_re, inv_im, bit0_re, mul_re, mul_zero, sub_self, neg_zero, algebra_map_eq_of_real], + ring, +end + +lemma inner_.conj_symm (x y : E) : conj (inner_ 𝕜 y x) = inner_ 𝕜 x y := +begin + simp only [inner_], + have h4 : conj (4⁻¹ : 𝕜) = 4⁻¹, + { rw [is_R_or_C.conj_inv, ←of_real_one, ←of_real_bit0, ←of_real_bit0, conj_of_real] }, + rw [map_mul, h4], + congr' 1, + simp only [map_sub, map_add, algebra_map_eq_of_real, ←of_real_mul, conj_of_real, map_mul, conj_I], + rw [add_comm y x, norm_sub_rev], + by_cases hI : (I : 𝕜) = 0, + { simp only [hI, neg_zero, zero_mul] }, + have h₁ : ‖(I : 𝕜) • y - x‖ = ‖(I : 𝕜) • x + y‖, + { transitivity ‖(I : 𝕜) • ((I : 𝕜) • y - x)‖, + { rw [norm_smul, norm_I_of_ne_zero hI, one_mul] }, + { rw [smul_sub, smul_smul, I_mul_I_of_nonzero hI, neg_one_smul, ←neg_add', add_comm, + norm_neg] } }, + have h₂ : ‖(I : 𝕜) • y + x‖ = ‖(I : 𝕜) • x - y‖, + { transitivity ‖(I : 𝕜) • ((I : 𝕜) • y + x)‖, + { rw [norm_smul, norm_I_of_ne_zero hI, one_mul] }, + { rw [smul_add, smul_smul, I_mul_I_of_nonzero hI, neg_one_smul, ←neg_add_eq_sub] }}, + rw [h₁, h₂, ←sub_add_eq_add_sub], + simp only [neg_mul, sub_eq_add_neg, neg_neg], +end + +variables [inner_product_spaceable E] + +private lemma add_left_aux1 (x y z : E) : + ‖x + y + z‖ * ‖x + y + z‖ = + (‖2 • x + y‖ * ‖2 • x + y‖ + ‖2 • z + y‖ * ‖2 • z + y‖) / 2 - ‖x - z‖ * ‖x - z‖ := +begin + rw [eq_sub_iff_add_eq, eq_div_iff (two_ne_zero' ℝ), mul_comm _ (2 : ℝ), eq_comm], + convert parallelogram_identity (x + y + z) (x - z) using 4; { rw two_smul, abel } +end + +private lemma add_left_aux2 (x y z : E) : + ‖x + y - z‖ * ‖x + y - z‖ = + (‖2 • x + y‖ * ‖2 • x + y‖ + ‖y - 2 • z‖ * ‖y - 2 • z‖) / 2 - ‖x + z‖ * ‖x + z‖ := +begin + rw [eq_sub_iff_add_eq, eq_div_iff (two_ne_zero' ℝ), mul_comm _ (2 : ℝ), eq_comm], + have h₀ := parallelogram_identity (x + y - z) (x + z), + convert h₀ using 4; { rw two_smul, abel } +end + +private lemma add_left_aux2' (x y z : E) : + ‖x + y + z‖ * ‖x + y + z‖ - ‖x + y - z‖ * ‖x + y - z‖ = + ‖x + z‖ * ‖x + z‖ - ‖x - z‖ * ‖x - z‖ + + (‖2 • z + y‖ * ‖2 • z + y‖ - ‖y - 2 • z‖ * ‖y - 2 • z‖) / 2 := +by { rw [add_left_aux1 , add_left_aux2], ring } + +private lemma add_left_aux3 (y z : E) : + ‖2 • z + y‖ * ‖2 • z + y‖ = 2 * (‖y + z‖ * ‖y + z‖ + ‖z‖ * ‖z‖) - ‖y‖ * ‖y‖ := +begin + apply eq_sub_of_add_eq, + convert parallelogram_identity (y + z) z using 4; { try { rw two_smul }, abel } +end + +private lemma add_left_aux4 (y z : E) : + ‖y - 2 • z‖ * ‖y - 2 • z‖ = 2 * (‖y - z‖ * ‖y - z‖ + ‖z‖ * ‖z‖) - ‖y‖ * ‖y‖ := +begin + apply eq_sub_of_add_eq', + have h₀ := parallelogram_identity (y - z) z, + convert h₀ using 4; { try { rw two_smul }, abel } +end + +private lemma add_left_aux4' (y z : E) : + (‖2 • z + y‖ * ‖2 • z + y‖ - ‖y - 2 • z‖ * ‖y - 2 • z‖) / 2 = + (‖y + z‖ * ‖y + z‖) - (‖y - z‖ * ‖y - z‖) := +by { rw [add_left_aux3 , add_left_aux4], ring } + +private lemma add_left_aux5 (x y z : E) : + ‖(I : 𝕜) • (x + y) + z‖ * ‖(I : 𝕜) • (x + y) + z‖ = + (‖(I : 𝕜) • (2 • x + y)‖ * ‖(I : 𝕜) • (2 • x + y)‖ + + ‖(I : 𝕜) • y + 2 • z‖ * ‖(I : 𝕜) • y + 2 • z‖) / 2 - ‖(I : 𝕜) • x - z‖ * ‖(I : 𝕜) • x - z‖ := +begin + rw [eq_sub_iff_add_eq, eq_div_iff (two_ne_zero' ℝ), mul_comm _ (2 : ℝ), eq_comm], + have h₀ := parallelogram_identity ((I : 𝕜) • (x + y) + z) ((I : 𝕜) • x - z), + convert h₀ using 4; { try { simp only [two_smul, smul_add] }, abel } +end + +private lemma add_left_aux6 (x y z : E) : + ‖(I : 𝕜) • (x + y) - z‖ * ‖(I : 𝕜) • (x + y) - z‖ = + (‖(I : 𝕜) • (2 • x + y)‖ * ‖(I : 𝕜) • (2 • x + y)‖ + + ‖(I : 𝕜) • y - 2 • z‖ * ‖(I : 𝕜) • y - 2 • z‖) / 2 - + ‖(I : 𝕜) • x + z‖ * ‖(I : 𝕜) • x + z‖ := +begin + rw [eq_sub_iff_add_eq, eq_div_iff (two_ne_zero' ℝ), mul_comm _ (2 : ℝ), eq_comm], + have h₀ := parallelogram_identity ((I : 𝕜) • (x + y) - z) ((I : 𝕜) • x + z), + convert h₀ using 4; { try { simp only [two_smul, smul_add] }, abel } +end + +private lemma add_left_aux7 (y z : E) : + ‖(I : 𝕜) • y + 2 • z‖ * ‖(I : 𝕜) • y + 2 • z‖ = + 2 * (‖(I : 𝕜) • y + z‖ * ‖(I : 𝕜) • y + z‖ + ‖z‖ * ‖z‖) - ‖(I : 𝕜) • y‖ * ‖(I : 𝕜) • y‖ := +begin + apply eq_sub_of_add_eq, + have h₀ := parallelogram_identity ((I : 𝕜) • y + z) z, + convert h₀ using 4; { try { simp only [two_smul, smul_add] }, abel } +end + +private lemma add_left_aux8 (y z : E) : + ‖(I : 𝕜) • y - 2 • z‖ * ‖(I : 𝕜) • y - 2 • z‖ = + 2 * (‖(I : 𝕜) • y - z‖ * ‖(I : 𝕜) • y - z‖ + ‖z‖ * ‖z‖) - ‖(I : 𝕜) • y‖ * ‖(I : 𝕜) • y‖ := +begin + apply eq_sub_of_add_eq', + have h₀ := parallelogram_identity ((I : 𝕜) • y - z) z, + convert h₀ using 4; { try { simp only [two_smul, smul_add] }, abel } +end + +lemma add_left (x y z : E) : inner_ 𝕜 (x + y) z = inner_ 𝕜 x z + inner_ 𝕜 y z := +begin + simp only [inner_, ←mul_add], + congr, + simp only [mul_assoc, ←map_mul, add_sub_assoc, ←mul_sub, ←map_sub], + rw add_add_add_comm, + simp only [←map_add, ←mul_add], + congr, + { rw [←add_sub_assoc, add_left_aux2', add_left_aux4'] }, + { rw [add_left_aux5, add_left_aux6, add_left_aux7, add_left_aux8], + simp only [map_sub, map_mul, map_add, div_eq_mul_inv], + ring } +end + +lemma nat (n : ℕ) (x y : E) : inner_ 𝕜 ((n : 𝕜) • x) y = (n : 𝕜) * inner_ 𝕜 x y := +begin + induction n with n ih, + { simp only [inner_, nat.nat_zero_eq_zero, zero_sub, nat.cast_zero, zero_mul, eq_self_iff_true, + zero_smul, zero_add, mul_zero, sub_self, norm_neg, smul_zero], }, + { simp only [nat.cast_succ, add_smul, one_smul], + rw [add_left, ih, add_mul, one_mul] } +end + +private lemma nat_prop (r : ℕ) : inner_prop E (r : 𝕜) := +λ x y, by { simp only [map_nat_cast], exact nat r x y } + +private lemma int_prop (n : ℤ) : inner_prop E (n : 𝕜) := +begin + intros x y, + rw ←n.sign_mul_nat_abs, + simp only [int.cast_coe_nat, map_nat_cast, map_int_cast, int.cast_mul, map_mul, mul_smul], + obtain hn | rfl | hn := lt_trichotomy n 0, + { rw [int.sign_eq_neg_one_of_neg hn, inner_prop_neg_one ((n.nat_abs : 𝕜) • x), nat], + simp only [map_neg, neg_mul, one_mul, mul_eq_mul_left_iff, true_or, + int.nat_abs_eq_zero, eq_self_iff_true, int.cast_one, map_one, neg_inj, nat.cast_eq_zero, + int.cast_neg] }, + { simp only [inner_, int.cast_zero, zero_sub, nat.cast_zero, zero_mul, eq_self_iff_true, + int.sign_zero, zero_smul, zero_add, mul_zero, smul_zero, sub_self, norm_neg, + int.nat_abs_zero] }, + { rw int.sign_eq_one_of_pos hn, + simp only [one_mul, mul_eq_mul_left_iff, true_or, int.nat_abs_eq_zero, eq_self_iff_true, + int.cast_one, one_smul, nat.cast_eq_zero, nat] } +end + +private lemma rat_prop (r : ℚ) : inner_prop E (r : 𝕜) := +begin + intros x y, + have : (r.denom : 𝕜) ≠ 0, + { haveI : char_zero 𝕜 := is_R_or_C.char_zero_R_or_C, + exact_mod_cast r.pos.ne' }, + rw [←r.num_div_denom, ←mul_right_inj' this, ←nat r.denom _ y, smul_smul, rat.cast_div], + simp only [map_nat_cast, rat.cast_coe_nat, map_int_cast, rat.cast_coe_int, map_div₀], + rw [←mul_assoc, mul_div_cancel' _ this, int_prop _ x, map_int_cast], +end + +private lemma real_prop (r : ℝ) : inner_prop E (r : 𝕜) := +begin + intros x y, + revert r, + rw ←function.funext_iff, + refine rat.dense_embedding_coe_real.dense.equalizer _ _ (funext $ λ X, _), + { exact (continuous_of_real.smul continuous_const).inner_ continuous_const }, + { exact (continuous_conj.comp continuous_of_real).mul continuous_const }, + { simp only [function.comp_app, is_R_or_C.of_real_rat_cast, rat_prop _ _] } +end + +private lemma I_prop : inner_prop E (I : 𝕜) := +begin + by_cases hI : (I : 𝕜) = 0, + { rw [hI, ←nat.cast_zero], exact nat_prop _ }, + intros x y, + have hI' : (-I : 𝕜) * I = 1, + { rw [←inv_I, inv_mul_cancel hI], }, + rw [conj_I, inner_, inner_, mul_left_comm], + congr' 1, + rw [smul_smul, I_mul_I_of_nonzero hI, neg_one_smul], + rw [mul_sub, mul_add, mul_sub, + mul_assoc I (𝓚 ‖I • x - y‖), ←mul_assoc (-I) I, hI', one_mul, + mul_assoc I (𝓚 ‖I • x + y‖), ←mul_assoc (-I) I, hI', one_mul], + have h₁ : ‖-x - y‖ = ‖x + y‖, + { rw [←neg_add', norm_neg], }, + have h₂ : ‖-x + y‖ = ‖x - y‖, + { rw [←neg_sub, norm_neg, sub_eq_neg_add], }, + rw [h₁, h₂], + simp only [sub_eq_add_neg, mul_assoc], + rw [←neg_mul_eq_neg_mul, ←neg_mul_eq_neg_mul], + abel +end + +lemma inner_prop (r : 𝕜) : inner_prop E r := +begin + intros x y, + rw [←re_add_im r, add_smul, add_left, real_prop _ x, ←smul_smul, real_prop _ _ y, I_prop, + map_add, map_mul, conj_of_real, conj_of_real, conj_I], + ring, +end + +end inner_product_spaceable + +open inner_product_spaceable + +/-- **Fréchet–von Neumann–Jordan Theorem**. A normed space `E` whose norm satisfies the +parallelogram identity can be given a compatible inner product. -/ +noncomputable def inner_product_space.of_norm + (h : ∀ x y : E, ‖x + y‖ * ‖x + y‖ + ‖x - y‖ * ‖x - y‖ = 2 * (‖x‖ * ‖x‖ + ‖y‖ * ‖y‖)) : + inner_product_space 𝕜 E := +begin + haveI : inner_product_spaceable E := ⟨h⟩, + exact + { inner := inner_ 𝕜, + norm_sq_eq_inner := inner_.norm_sq, + conj_symm := inner_.conj_symm, + add_left := add_left, + smul_left := λ _ _ _, inner_prop _ _ _ } +end + +variables (𝕜 E) [inner_product_spaceable E] + +/-- **Fréchet–von Neumann–Jordan Theorem**. A normed space `E` whose norm satisfies the +parallelogram identity can be given a compatible inner product. Do +`casesI nonempty_inner_product_space 𝕜 E` to locally upgrade `inner_product_spaceable E` to +`inner_product_space 𝕜 E`. -/ +lemma nonempty_inner_product_space : nonempty (inner_product_space 𝕜 E) := +⟨{ inner := inner_ 𝕜, + norm_sq_eq_inner := inner_.norm_sq, + conj_symm := inner_.conj_symm, + add_left := add_left, + smul_left := λ _ _ _, inner_prop _ _ _ }⟩ + +variables {𝕜 E} [normed_space ℝ E] + +-- TODO: Replace `inner_product_space.to_uniform_convex_space` +@[priority 100] -- See note [lower instance priority] +instance inner_product_spaceable.to_uniform_convex_space : uniform_convex_space E := +by { casesI nonempty_inner_product_space ℝ E, apply_instance } diff --git a/src/data/is_R_or_C/basic.lean b/src/data/is_R_or_C/basic.lean index c747fef34bf60..71714a3f5cf1c 100644 --- a/src/data/is_R_or_C/basic.lean +++ b/src/data/is_R_or_C/basic.lean @@ -491,6 +491,12 @@ begin exact div_le_one_of_le (abs_im_le_norm _) (norm_nonneg _) end +lemma norm_I_of_ne_zero (hI : (I : K) ≠ 0) : ‖(I : K)‖ = 1 := +begin + rw [← mul_self_inj_of_nonneg (norm_nonneg I) zero_le_one, one_mul, ← norm_mul, + I_mul_I_of_nonzero hI, norm_neg, norm_one], +end + lemma re_eq_norm_of_mul_conj (x : K) : re (x * conj x) = ‖x * conj x‖ := by rw [mul_conj, of_real_re, norm_of_real, abs_of_nonneg (norm_sq_nonneg _)] From cd391184c85986113f8c00844cfe6dda1d34be3d Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 13 Jul 2023 10:41:08 +0000 Subject: [PATCH 1248/1291] feat(*/prod): `prod_prod_prod` equivs (#19235) These send `((a, b), (c, d))` to `((a, c), (b, d))`, and this commit provides this bundled as `equiv`, `add_equiv`, `mul_equiv`, `ring_equiv`, and `linear_equiv`. We already have something analogous for `tensor_product`. --- src/algebra/group/prod.lean | 20 ++++++++++++++++++++ src/algebra/ring/prod.lean | 29 ++++++++++++++++++++++++++++- src/linear_algebra/prod.lean | 24 +++++++++++++++++++++++- src/logic/equiv/basic.lean | 11 +++++++++++ 4 files changed, 82 insertions(+), 2 deletions(-) diff --git a/src/algebra/group/prod.lean b/src/algebra/group/prod.lean index 3fabcc474d8b4..2546fc56e9435 100644 --- a/src/algebra/group/prod.lean +++ b/src/algebra/group/prod.lean @@ -465,6 +465,26 @@ def prod_comm : M × N ≃* N × M := variables {M' N' : Type*} [mul_one_class M'] [mul_one_class N'] +section +variables (M N M' N') + +/-- Four-way commutativity of `prod`. The name matches `mul_mul_mul_comm`. -/ +@[to_additive prod_prod_prod_comm "Four-way commutativity of `prod`. +The name matches `mul_mul_mul_comm`", simps apply] +def prod_prod_prod_comm : (M × N) × (M' × N') ≃* (M × M') × (N × N') := +{ to_fun := λ mnmn, ((mnmn.1.1, mnmn.2.1), (mnmn.1.2, mnmn.2.2)), + inv_fun := λ mmnn, ((mmnn.1.1, mmnn.2.1), (mmnn.1.2, mmnn.2.2)), + map_mul' := λ mnmn mnmn', rfl, + ..equiv.prod_prod_prod_comm M N M' N', } + +@[simp, to_additive] lemma prod_prod_prod_comm_to_equiv : + (prod_prod_prod_comm M N M' N').to_equiv = equiv.prod_prod_prod_comm M N M' N' := rfl + +@[simp] lemma prod_prod_prod_comm_symm : + (prod_prod_prod_comm M N M' N').symm = prod_prod_prod_comm M M' N N' := rfl + +end + /--Product of multiplicative isomorphisms; the maps come from `equiv.prod_congr`.-/ @[to_additive prod_congr "Product of additive isomorphisms; the maps come from `equiv.prod_congr`."] def prod_congr (f : M ≃* M') (g : N ≃* N') : M × N ≃* M' × N' := diff --git a/src/algebra/ring/prod.lean b/src/algebra/ring/prod.lean index 1ac110c8c81a5..5f5c2da349c5e 100644 --- a/src/algebra/ring/prod.lean +++ b/src/algebra/ring/prod.lean @@ -212,7 +212,9 @@ end prod_map end ring_hom namespace ring_equiv -variables {R S} [non_assoc_semiring R] [non_assoc_semiring S] +variables {R S R' S'} +variables [non_assoc_semiring R] [non_assoc_semiring S] +variables [non_assoc_semiring R'] [non_assoc_semiring S'] /-- Swapping components as an equivalence of (semi)rings. -/ def prod_comm : R × S ≃+* S × R := @@ -229,6 +231,31 @@ ring_hom.ext $ λ _, rfl (ring_hom.snd S R).comp ↑(prod_comm : R × S ≃+* S × R) = ring_hom.fst R S := ring_hom.ext $ λ _, rfl +section +variables (R R' S S') + +/-- Four-way commutativity of `prod`. The name matches `mul_mul_mul_comm`. -/ +@[simps apply] +def prod_prod_prod_comm : (R × R') × (S × S') ≃+* (R × S) × (R' × S') := +{ to_fun := λ rrss, ((rrss.1.1, rrss.2.1), (rrss.1.2, rrss.2.2)), + inv_fun := λ rsrs, ((rsrs.1.1, rsrs.2.1), (rsrs.1.2, rsrs.2.2)), + .. add_equiv.prod_prod_prod_comm R R' S S', + .. mul_equiv.prod_prod_prod_comm R R' S S' } + +@[simp] lemma prod_prod_prod_comm_symm : + (prod_prod_prod_comm R R' S S').symm = prod_prod_prod_comm R S R' S' := rfl + +@[simp] lemma prod_prod_prod_comm_to_add_equiv : + (prod_prod_prod_comm R R' S S').to_add_equiv = add_equiv.prod_prod_prod_comm R R' S S' := rfl + +@[simp] lemma prod_prod_prod_comm_to_mul_equiv : + (prod_prod_prod_comm R R' S S').to_mul_equiv = mul_equiv.prod_prod_prod_comm R R' S S' := rfl + +@[simp] lemma prod_prod_prod_comm_to_equiv : + (prod_prod_prod_comm R R' S S').to_equiv = equiv.prod_prod_prod_comm R R' S S' := rfl + +end + variables (R S) [subsingleton S] /-- A ring `R` is isomorphic to `R × S` when `S` is the zero ring -/ diff --git a/src/linear_algebra/prod.lean b/src/linear_algebra/prod.lean index 2b190f6abffca..546d089982d5b 100644 --- a/src/linear_algebra/prod.lean +++ b/src/linear_algebra/prod.lean @@ -227,7 +227,7 @@ def prod_map (f : M →ₗ[R] M₃) (g : M₂ →ₗ[R] M₄) : (M × M₂) → lemma coe_prod_map (f : M →ₗ[R] M₃) (g : M₂ →ₗ[R] M₄) : ⇑(f.prod_map g) = prod.map f g := rfl - + @[simp] theorem prod_map_apply (f : M →ₗ[R] M₃) (g : M₂ →ₗ[R] M₄) (x) : f.prod_map g x = (f x.1, g x.2) := rfl @@ -559,6 +559,28 @@ def prod_comm (R M N : Type*) [semiring R] [add_comm_monoid M] [add_comm_monoid map_smul' := λ r ⟨m, n⟩, rfl, ..add_equiv.prod_comm } +section +variables (R M M₂ M₃ M₄) +variables [semiring R] +variables [add_comm_monoid M] [add_comm_monoid M₂] [add_comm_monoid M₃] [add_comm_monoid M₄] +variables [module R M] [module R M₂] [module R M₃] [module R M₄] + +/-- Four-way commutativity of `prod`. The name matches `mul_mul_mul_comm`. -/ +@[simps apply] +def prod_prod_prod_comm : ((M × M₂) × (M₃ × M₄)) ≃ₗ[R] (M × M₃) × (M₂ × M₄) := +{ to_fun := λ mnmn, ((mnmn.1.1, mnmn.2.1), (mnmn.1.2, mnmn.2.2)), + inv_fun := λ mmnn, ((mmnn.1.1, mmnn.2.1), (mmnn.1.2, mmnn.2.2)), + map_smul' := λ c mnmn, rfl, + ..add_equiv.prod_prod_prod_comm M M₂ M₃ M₄ } + +@[simp] lemma prod_prod_prod_comm_symm : + (prod_prod_prod_comm R M M₂ M₃ M₄).symm = prod_prod_prod_comm R M M₃ M₂ M₄ := rfl + +@[simp] lemma prod_prod_prod_comm_to_add_equiv : + (prod_prod_prod_comm R M M₂ M₃ M₄).to_add_equiv = add_equiv.prod_prod_prod_comm M M₂ M₃ M₄ := rfl + +end + section variables [semiring R] diff --git a/src/logic/equiv/basic.lean b/src/logic/equiv/basic.lean index 6b6601d3d857c..84745166d6e14 100644 --- a/src/logic/equiv/basic.lean +++ b/src/logic/equiv/basic.lean @@ -105,6 +105,17 @@ def prod_comm (α β : Type*) : α × β ≃ β × α := @[simps] def prod_assoc (α β γ : Sort*) : (α × β) × γ ≃ α × (β × γ) := ⟨λ p, (p.1.1, p.1.2, p.2), λ p, ((p.1, p.2.1), p.2.2), λ ⟨⟨a, b⟩, c⟩, rfl, λ ⟨a, ⟨b, c⟩⟩, rfl⟩ +/-- Four-way commutativity of `prod`. The name matches `mul_mul_mul_comm`. -/ +@[simps apply] +def prod_prod_prod_comm (α β γ δ : Type*) : (α × β) × (γ × δ) ≃ (α × γ) × (β × δ) := +{ to_fun := λ abcd, ((abcd.1.1, abcd.2.1), (abcd.1.2, abcd.2.2)), + inv_fun := λ acbd, ((acbd.1.1, acbd.2.1), (acbd.1.2, acbd.2.2)), + left_inv := λ ⟨⟨a, b⟩, ⟨c, d⟩⟩, rfl, + right_inv := λ ⟨⟨a, c⟩, ⟨b, d⟩⟩, rfl, } + +@[simp] lemma prod_prod_prod_comm_symm (α β γ δ : Type*) : + (prod_prod_prod_comm α β γ δ).symm = prod_prod_prod_comm α γ β δ := rfl + /-- Functions on `α × β` are equivalent to functions `α → β → γ`. -/ @[simps {fully_applied := ff}] def curry (α β γ : Type*) : (α × β → γ) ≃ (α → β → γ) := From d608fc5d4e69d4cc21885913fb573a88b0deb521 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 13 Jul 2023 13:57:44 +0000 Subject: [PATCH 1249/1291] feat(analysis/calculus/fderiv/mul): derivative of inverse in division rings (#19127) --- src/analysis/calculus/fderiv/mul.lean | 90 ++++++++++++++++++++++++- src/analysis/normed_space/spectrum.lean | 2 +- 2 files changed, 89 insertions(+), 3 deletions(-) diff --git a/src/analysis/calculus/fderiv/mul.lean b/src/analysis/calculus/fderiv/mul.lean index f52bc203dcd4d..12ccaad983130 100644 --- a/src/analysis/calculus/fderiv/mul.lean +++ b/src/analysis/calculus/fderiv/mul.lean @@ -456,13 +456,99 @@ begin units.inv_mul, add_sub_cancel'_right, mul_sub, sub_mul, one_mul, sub_neg_eq_add] end -lemma differentiable_at_inverse (x : Rˣ) : differentiable_at 𝕜 (@ring.inverse R _) x := -(has_fderiv_at_ring_inverse x).differentiable_at +lemma differentiable_at_inverse {x : R} (hx : is_unit x) : + differentiable_at 𝕜 (@ring.inverse R _) x := +let ⟨u, hu⟩ := hx in hu ▸ (has_fderiv_at_ring_inverse u).differentiable_at + +lemma differentiable_within_at_inverse {x : R} (hx : is_unit x) (s : set R): + differentiable_within_at 𝕜 (@ring.inverse R _) s x := +(differentiable_at_inverse hx).differentiable_within_at + +lemma differentiable_on_inverse : differentiable_on 𝕜 (@ring.inverse R _) {x | is_unit x} := +λ x hx, differentiable_within_at_inverse hx _ lemma fderiv_inverse (x : Rˣ) : fderiv 𝕜 (@ring.inverse R _) x = - mul_left_right 𝕜 R ↑x⁻¹ ↑x⁻¹ := (has_fderiv_at_ring_inverse x).fderiv +variables {h : E → R} {z : E} {S : set E} + +lemma differentiable_within_at.inverse (hf : differentiable_within_at 𝕜 h S z) + (hz : is_unit (h z)) : + differentiable_within_at 𝕜 (λ x, ring.inverse (h x)) S z := +(differentiable_at_inverse hz).comp_differentiable_within_at z hf + +@[simp] lemma differentiable_at.inverse (hf : differentiable_at 𝕜 h z) (hz : is_unit (h z)) : + differentiable_at 𝕜 (λ x, ring.inverse (h x)) z := +(differentiable_at_inverse hz).comp z hf + +lemma differentiable_on.inverse (hf : differentiable_on 𝕜 h S) (hz : ∀ x ∈ S, is_unit (h x)) : + differentiable_on 𝕜 (λ x, ring.inverse (h x)) S := +λ x h, (hf x h).inverse (hz x h) + +@[simp] lemma differentiable.inverse (hf : differentiable 𝕜 h) (hz : ∀ x, is_unit (h x)) : + differentiable 𝕜 (λ x, ring.inverse (h x)) := +λ x, (hf x).inverse (hz x) + end algebra_inverse +/-! ### Derivative of the inverse in a division ring + +Note these lemmas are primed as they need `complete_space R`, whereas the other lemmas in +`deriv/inv.lean` do not, but instead need `nontrivially_normed_field R`. +-/ + +section division_ring_inverse +variables {R : Type*} [normed_division_ring R] [normed_algebra 𝕜 R] [complete_space R] +open normed_ring continuous_linear_map ring + +/-- At an invertible element `x` of a normed division algebra `R`, the Fréchet derivative of the +inversion operation is the linear map `λ t, - x⁻¹ * t * x⁻¹`. -/ +lemma has_fderiv_at_inv' {x : R} (hx : x ≠ 0) : + has_fderiv_at has_inv.inv (-mul_left_right 𝕜 R x⁻¹ x⁻¹) x := +by simpa using has_fderiv_at_ring_inverse (units.mk0 _ hx) + +lemma differentiable_at_inv' {x : R} (hx : x ≠ 0) : differentiable_at 𝕜 has_inv.inv x := +(has_fderiv_at_inv' hx).differentiable_at + +lemma differentiable_within_at_inv' {x : R} (hx : x ≠ 0) (s : set R): + differentiable_within_at 𝕜 (λx, x⁻¹) s x := +(differentiable_at_inv' hx).differentiable_within_at + +lemma differentiable_on_inv' : differentiable_on 𝕜 (λ x : R, x⁻¹) {x | x ≠ 0} := +λ x hx, differentiable_within_at_inv' hx _ + +/-- Non-commutative version of `fderiv_inv` -/ +lemma fderiv_inv' {x : R} (hx : x ≠ 0) : + fderiv 𝕜 has_inv.inv x = - mul_left_right 𝕜 R x⁻¹ x⁻¹ := +(has_fderiv_at_inv' hx).fderiv + +/-- Non-commutative version of `fderiv_within_inv` -/ +lemma fderiv_within_inv' {s : set R} {x : R} (hx : x ≠ 0) (hxs : unique_diff_within_at 𝕜 s x) : + fderiv_within 𝕜 (λ x, x⁻¹) s x = - mul_left_right 𝕜 R x⁻¹ x⁻¹ := +begin + rw differentiable_at.fderiv_within (differentiable_at_inv' hx) hxs, + exact fderiv_inv' hx +end + +variables {h : E → R} {z : E} {S : set E} + +lemma differentiable_within_at.inv' (hf : differentiable_within_at 𝕜 h S z) (hz : h z ≠ 0) : + differentiable_within_at 𝕜 (λ x, (h x)⁻¹) S z := +(differentiable_at_inv' hz).comp_differentiable_within_at z hf + +@[simp] lemma differentiable_at.inv' (hf : differentiable_at 𝕜 h z) (hz : h z ≠ 0) : + differentiable_at 𝕜 (λ x, (h x)⁻¹) z := +(differentiable_at_inv' hz).comp z hf + +lemma differentiable_on.inv' (hf : differentiable_on 𝕜 h S) (hz : ∀ x ∈ S, h x ≠ 0) : + differentiable_on 𝕜 (λ x, (h x)⁻¹) S := +λ x h, (hf x h).inv' (hz x h) + +@[simp] lemma differentiable.inv' (hf : differentiable 𝕜 h) (hz : ∀ x, h x ≠ 0) : + differentiable 𝕜 (λ x, (h x)⁻¹) := +λ x, (hf x).inv' (hz x) + +end division_ring_inverse + end diff --git a/src/analysis/normed_space/spectrum.lean b/src/analysis/normed_space/spectrum.lean index 09dbf5e2d5006..d222373613d9b 100644 --- a/src/analysis/normed_space/spectrum.lean +++ b/src/analysis/normed_space/spectrum.lean @@ -314,7 +314,7 @@ begin simpa only [norm_to_nnreal, real.to_nnreal_coe] using real.to_nnreal_mono (mem_closed_ball_zero_iff.mp z_mem) }, have H₁ : differentiable 𝕜 (λ w : 𝕜, 1 - w • a) := (differentiable_id.smul_const a).const_sub 1, - exact differentiable_at.comp z (differentiable_at_inverse hu.unit) (H₁.differentiable_at), + exact differentiable_at.comp z (differentiable_at_inverse hu) (H₁.differentiable_at), end end one_sub_smul From d07245fd37786daa997af4f1a73a49fa3b748408 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 14 Jul 2023 16:30:09 +0000 Subject: [PATCH 1250/1291] =?UTF-8?q?feat(group=5Ftheory/order=5Fof=5Felem?= =?UTF-8?q?ent):=20Order=20in=20`=CE=B1=20=C3=97=20=CE=B2`=20(#18719)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The order of `(a, b)` is the lcm of the orders of `a` and `b`. Match `pow` and `zpow` lemmas. Also some `variables` noise because I could not use `x` to mean what I wanted, and incidentally the type `A` was mostly unused. --- src/data/prod/basic.lean | 2 +- src/dynamics/periodic_pts.lean | 25 +++++++ src/group_theory/order_of_element.lean | 90 ++++++++++++++++++++------ 3 files changed, 95 insertions(+), 22 deletions(-) diff --git a/src/data/prod/basic.lean b/src/data/prod/basic.lean index a2a159aced62e..eb8f6034c7058 100644 --- a/src/data/prod/basic.lean +++ b/src/data/prod/basic.lean @@ -97,7 +97,7 @@ funext (λ p, ext (map_fst f g p) (map_snd f g p)) lemma id_prod : (λ (p : α × β), (p.1, p.2)) = id := funext $ λ ⟨a, b⟩, rfl -lemma map_id : (prod.map (@id α) (@id β)) = id := +@[simp] lemma map_id : (prod.map (@id α) (@id β)) = id := id_prod lemma fst_surjective [h : nonempty β] : function.surjective (@fst α β) := diff --git a/src/dynamics/periodic_pts.lean b/src/dynamics/periodic_pts.lean index 30e6211e65126..353804e7661ea 100644 --- a/src/dynamics/periodic_pts.lean +++ b/src/dynamics/periodic_pts.lean @@ -519,6 +519,31 @@ end end function +namespace function +variables {α β : Type*} {f : α → α} {g : β → β} {x : α × β} {a : α} {b : β} {m n : ℕ} + +@[simp] lemma iterate_prod_map (f : α → α) (g : β → β) (n : ℕ) : + (prod.map f g)^[n] = prod.map (f^[n]) (g^[n]) := by induction n; simp [*, prod.map_comp_map] + +@[simp] lemma is_fixed_pt_prod_map (x : α × β) : + is_fixed_pt (prod.map f g) x ↔ is_fixed_pt f x.1 ∧ is_fixed_pt g x.2 := prod.ext_iff + +@[simp] lemma is_periodic_pt_prod_map (x : α × β) : + is_periodic_pt (prod.map f g) n x ↔ is_periodic_pt f n x.1 ∧ is_periodic_pt g n x.2 := +by simp [is_periodic_pt] + +lemma minimal_period_prod_map (f : α → α) (g : β → β) (x : α × β) : + minimal_period (prod.map f g) x = (minimal_period f x.1).lcm (minimal_period g x.2) := +eq_of_forall_dvd $ by cases x; simp [←is_periodic_pt_iff_minimal_period_dvd, nat.lcm_dvd_iff] + +lemma minimal_period_fst_dvd : minimal_period f x.1 ∣ minimal_period (prod.map f g) x := +by { rw minimal_period_prod_map, exact nat.dvd_lcm_left _ _ } + +lemma minimal_period_snd_dvd : minimal_period g x.2 ∣ minimal_period (prod.map f g) x := +by { rw minimal_period_prod_map, exact nat.dvd_lcm_right _ _ } + +end function + namespace mul_action open function diff --git a/src/group_theory/order_of_element.lean b/src/group_theory/order_of_element.lean index 44eb664b3460a..2f1065af425db 100644 --- a/src/group_theory/order_of_element.lean +++ b/src/group_theory/order_of_element.lean @@ -3,8 +3,9 @@ Copyright (c) 2018 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Julian Kuelshammer -/ +import algebra.gcd_monoid.finset import algebra.hom.iterate -import data.nat.modeq +import data.int.modeq import data.set.pointwise.basic import data.set.intervals.infinite import dynamics.periodic_pts @@ -35,14 +36,11 @@ order of an element open function nat open_locale pointwise -universes u v - -variables {G : Type u} {A : Type v} -variables {x y : G} {a b : A} {n m : ℕ} +variables {G H A α β : Type*} section monoid_add_monoid -variables [monoid G] [add_monoid A] +variables [monoid G] [add_monoid A] {x y : G} {a b : A} {n m : ℕ} section is_of_fin_order @@ -95,8 +93,7 @@ by { rw [is_of_fin_order_iff_pow_eq_one, is_of_fin_order_iff_pow_eq_one], norm_c /-- The image of an element of finite order has finite order. -/ @[to_additive add_monoid_hom.is_of_fin_order "The image of an element of finite additive order has finite additive order."] -lemma monoid_hom.is_of_fin_order - {H : Type v} [monoid H] (f : G →* H) {x : G} (h : is_of_fin_order x) : +lemma monoid_hom.is_of_fin_order [monoid H] (f : G →* H) {x : G} (h : is_of_fin_order x) : is_of_fin_order $ f x := (is_of_fin_order_iff_pow_eq_one _).mpr $ begin rcases (is_of_fin_order_iff_pow_eq_one _).mp h with ⟨n, npos, hn⟩, @@ -171,6 +168,11 @@ end lemma order_of_pos_iff : 0 < order_of x ↔ is_of_fin_order x := by rwa [iff_not_comm.mp order_of_eq_zero_iff, pos_iff_ne_zero] +@[to_additive is_of_fin_add_order.mono] +lemma is_of_fin_order.mono [monoid β] {y : β} (hx : is_of_fin_order x) + (h : order_of y ∣ order_of x) : is_of_fin_order y := +by { rw ←order_of_pos_iff at ⊢ hx, exact nat.pos_of_dvd_of_pos h hx } + @[to_additive nsmul_ne_zero_of_lt_add_order_of'] lemma pow_ne_one_of_lt_order_of' (n0 : n ≠ 0) (h : n < order_of x) : x ^ n ≠ 1 := λ j, not_is_periodic_pt_of_pos_of_lt_minimal_period n0 h @@ -311,11 +313,11 @@ end lemma order_of_dvd_lcm_mul : order_of y ∣ nat.lcm (order_of x) (order_of (x * y)) := begin by_cases h0 : order_of x = 0, - { rw [h0, lcm_zero_left], apply dvd_zero }, + { rw [h0, nat.lcm_zero_left], apply dvd_zero }, conv_lhs { rw [← one_mul y, ← pow_order_of_eq_one x, ← succ_pred_eq_of_pos (nat.pos_of_ne_zero h0), pow_succ', mul_assoc] }, exact (((commute.refl x).mul_right h).pow_left _).order_of_mul_dvd_lcm.trans - (lcm_dvd_iff.2 ⟨trans (order_of_pow_dvd _) (dvd_lcm_left _ _), dvd_lcm_right _ _⟩), + (nat.lcm_dvd_iff.2 ⟨trans (order_of_pow_dvd _) (dvd_lcm_left _ _), dvd_lcm_right _ _⟩), end @[to_additive add_order_of_add_dvd_mul_add_order_of] @@ -394,7 +396,7 @@ end p_prime end monoid_add_monoid section cancel_monoid -variables [left_cancel_monoid G] (x y) +variables [left_cancel_monoid G] (x y : G) {m n : ℕ} @[to_additive nsmul_injective_of_lt_add_order_of] lemma pow_injective_of_lt_order_of @@ -451,7 +453,7 @@ end end cancel_monoid section group -variables [group G] [add_group A] {x a} {i : ℤ} +variables [group G] {x y : G} {i : ℤ} /-- Inverses of elements of finite order have finite order. -/ @[to_additive "Inverses of elements of finite additive order have finite additive order."] @@ -560,7 +562,7 @@ end group section comm_monoid -variables [comm_monoid G] +variables [comm_monoid G] {x y : G} /-- Elements of finite order are closed under multiplication. -/ @[to_additive "Elements of finite additive order are closed under addition."] @@ -571,7 +573,7 @@ lemma is_of_fin_order.mul (hx : is_of_fin_order x) (hy : is_of_fin_order y) : end comm_monoid section finite_monoid -variables [monoid G] +variables [monoid G] {n : ℕ} open_locale big_operators @[to_additive sum_card_add_order_of_eq_card_nsmul_eq_zero] @@ -592,7 +594,7 @@ end finite_monoid section finite_cancel_monoid -- TODO: Of course everything also works for right_cancel_monoids. -variables [left_cancel_monoid G] [add_left_cancel_monoid A] +variables [left_cancel_monoid G] {x y : G} {n : ℕ} -- TODO: Use this to show that a finite left cancellative monoid is a group. @[to_additive] @@ -682,7 +684,7 @@ lemma order_eq_card_powers [fintype G] : order_of x = fintype.card (submonoid.po end finite_cancel_monoid section finite_group -variables [group G] [add_group A] +variables [group G] {x y : G} {n : ℕ} @[to_additive] lemma exists_zpow_eq_one [finite G] (x : G) : ∃ (i : ℤ) (H : i ≠ 0), x ^ (i : ℤ) = 1 := @@ -712,6 +714,23 @@ lemma mem_zpowers_iff_mem_range_order_of [finite G] [decidable_eq G] : y ∈ subgroup.zpowers x ↔ y ∈ (finset.range (order_of x)).image ((^) x : ℕ → G) := by rw [← mem_powers_iff_mem_zpowers, mem_powers_iff_mem_range_order_of] +@[to_additive] lemma zpow_eq_one_iff_modeq {n : ℤ} : x ^ n = 1 ↔ n ≡ 0 [ZMOD (order_of x)] := +by rw [int.modeq_zero_iff_dvd, order_of_dvd_iff_zpow_eq_one] + +@[to_additive] lemma zpow_eq_zpow_iff_modeq {m n : ℤ} : x ^ m = x ^ n ↔ m ≡ n [ZMOD (order_of x)] := +by rw [←mul_inv_eq_one, ←zpow_sub, zpow_eq_one_iff_modeq, int.modeq_iff_dvd, int.modeq_iff_dvd, + zero_sub, neg_sub] + +@[simp, to_additive] lemma injective_zpow_iff_not_is_of_fin_order : + injective (λ n : ℤ, x ^ n) ↔ ¬ is_of_fin_order x := +begin + refine ⟨_, λ h n m hnm, _⟩, + { simp_rw is_of_fin_order_iff_pow_eq_one, + rintro h ⟨n, hn, hx⟩, + exact nat.cast_ne_zero.2 hn.ne' (h $ by simpa using hx) }, + rwa [zpow_eq_zpow_iff_modeq, order_of_eq_zero_iff.2 h, nat.cast_zero, int.modeq_zero_iff] at hnm, +end + @[to_additive decidable_zmultiples] noncomputable instance decidable_zpowers : decidable_pred (∈ subgroup.zpowers x) := classical.dec_pred _ @@ -755,8 +774,8 @@ end variables [fintype G] -/-- See also `order_eq_card_zpowers'`. -/ -@[to_additive add_order_eq_card_zmultiples "See also `add_order_eq_card_zmultiples'`."] +/-- See also `nat.card_zpowers'`. -/ +@[to_additive add_order_eq_card_zmultiples "See also `nat.card_zmultiples`."] lemma order_eq_card_zpowers : order_of x = fintype.card (zpowers x) := (fintype.card_fin (order_of x)).symm.trans (fintype.card_eq.2 ⟨fin_equiv_zpowers x⟩) @@ -846,8 +865,6 @@ begin (congr_arg (∣ fintype.card K) (order_of_subgroup ⟨x, hx.2⟩)).mpr order_of_dvd_card_univ⟩, end -variable (a) - /-- TODO: Generalise to `submonoid.powers`.-/ @[to_additive image_range_add_order_of, nolint to_additive_doc] lemma image_range_order_of [decidable_eq G] : @@ -909,7 +926,7 @@ end pow_is_subgroup section linear_ordered_ring -variable [linear_ordered_ring G] +variables [linear_ordered_ring G] {x : G} lemma order_of_abs_ne_one (h : |x| ≠ 1) : order_of x = 0 := begin @@ -931,3 +948,34 @@ begin end end linear_ordered_ring + +section prod +variables [monoid α] [monoid β] {x : α × β} {a : α} {b : β} + +@[to_additive prod.add_order_of] protected lemma prod.order_of (x : α × β) : + order_of x = (order_of x.1).lcm (order_of x.2) := +minimal_period_prod_map _ _ _ + +@[to_additive add_order_of_fst_dvd_add_order_of] lemma order_of_fst_dvd_order_of : + order_of x.1 ∣ order_of x := +minimal_period_fst_dvd + +@[to_additive add_order_of_snd_dvd_add_order_of] lemma order_of_snd_dvd_order_of : + order_of x.2 ∣ order_of x := +minimal_period_snd_dvd + +@[to_additive is_of_fin_add_order.fst] +lemma is_of_fin_order.fst {x : α × β} (hx : is_of_fin_order x) : is_of_fin_order x.1 := +hx.mono order_of_fst_dvd_order_of + +@[to_additive is_of_fin_add_order.snd] +lemma is_of_fin_order.snd {x : α × β} (hx : is_of_fin_order x) : is_of_fin_order x.2 := +hx.mono order_of_snd_dvd_order_of + +@[to_additive is_of_fin_add_order.prod_mk] +lemma is_of_fin_order.prod_mk : is_of_fin_order a → is_of_fin_order b → is_of_fin_order (a, b) := +by simpa only [←order_of_pos_iff, prod.order_of] using nat.lcm_pos + +end prod + +-- TODO: Corresponding `pi` lemmas. We cannot currently state them here because of import cycles From 2c84c2c5496117349007d97104e7bbb471381592 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 14 Jul 2023 19:58:48 +0000 Subject: [PATCH 1251/1291] feat(order/well_founded_set): `prod.lex` is well-founded (#18665) Co-authored-by: Junyan Xu Co-authored-by: Eric Wieser --- src/order/well_founded.lean | 27 +++++----- src/order/well_founded_set.lean | 96 ++++++++++++++++++++++++++++++++- 2 files changed, 108 insertions(+), 15 deletions(-) diff --git a/src/order/well_founded.lean b/src/order/well_founded.lean index fb82de79a2a55..40aaca54853a1 100644 --- a/src/order/well_founded.lean +++ b/src/order/well_founded.lean @@ -21,21 +21,24 @@ and provide a few new definitions: `well_founded.min`, `well_founded.sup`, and ` and an induction principle `well_founded.induction_bot`. -/ -variables {α : Type*} +variables {α β γ : Type*} namespace well_founded +variables {r r' : α → α → Prop} -protected theorem is_asymm {α : Sort*} {r : α → α → Prop} (h : well_founded r) : is_asymm α r := -⟨h.asymmetric⟩ +protected theorem is_asymm (h : well_founded r) : is_asymm α r := ⟨h.asymmetric⟩ -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 := +protected theorem is_irrefl (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 +instance [has_well_founded α] : is_asymm α has_well_founded.r := has_well_founded.wf.is_asymm +instance [has_well_founded α] : is_irrefl α has_well_founded.r := is_asymm.is_irrefl + +lemma mono (hr : well_founded r) (h : ∀ a b, r' a b → r a b) : well_founded r' := +subrelation.wf h hr + +lemma on_fun {α β : Sort*} {r : β → β → Prop} {f : α → β} : + well_founded r → well_founded (r on f) := inv_image.wf _ /-- If `r` is a well-founded relation, then any nonempty set has a minimal element with respect to `r`. -/ @@ -110,8 +113,7 @@ end section linear_order -variables {β : Type*} [linear_order β] (h : well_founded ((<) : β → β → Prop)) - {γ : Type*} [partial_order γ] +variables [linear_order β] (h : well_founded ((<) : β → β → Prop)) [partial_order γ] theorem min_le {x : β} {s : set β} (hx : x ∈ s) (hne : s.nonempty := ⟨x, hx⟩) : h.min s hne ≤ x := @@ -149,8 +151,7 @@ end linear_order end well_founded namespace function - -variables {β : Type*} (f : α → β) +variables (f : α → β) section has_lt diff --git a/src/order/well_founded_set.lean b/src/order/well_founded_set.lean index 37a675eb498cb..d3c122a549c21 100644 --- a/src/order/well_founded_set.lean +++ b/src/order/well_founded_set.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Aaron Anderson. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Aaron Anderson -/ +import data.sigma.lex import order.antichain import order.order_iso_nat import order.well_founded @@ -45,7 +46,7 @@ Prove that `s` is partial well ordered iff it has no infinite descending chain o * [Nash-Williams, *On Well-Quasi-Ordering Finite Trees*][Nash-Williams63] -/ -variables {ι α β : Type*} +variables {ι α β γ : Type*} {π : ι → Type*} namespace set @@ -61,7 +62,7 @@ section well_founded_on variables {r r' : α → α → Prop} section any_rel -variables {s t : set α} {x y : α} +variables {f : β → α} {s t : set α} {x y : α} lemma well_founded_on_iff : s.well_founded_on r ↔ well_founded (λ (a b : α), r a b ∧ a ∈ s ∧ b ∈ s) := @@ -79,6 +80,26 @@ begin exact ⟨m, mt, λ x xt ⟨xm, xs, ms⟩, hst ⟨m, ⟨ms, mt⟩⟩⟩ } end +@[simp] lemma well_founded_on_univ : (univ : set α).well_founded_on r ↔ well_founded r := +by simp [well_founded_on_iff] + +lemma _root_.well_founded.well_founded_on : well_founded r → s.well_founded_on r := inv_image.wf _ + +@[simp] lemma well_founded_on_range : (range f).well_founded_on r ↔ well_founded (r on f) := +begin + let f' : β → range f := λ c, ⟨f c, c, rfl⟩, + refine ⟨λ h, (inv_image.wf f' h).mono $ λ c c', id, λ h, ⟨_⟩⟩, + rintro ⟨_, c, rfl⟩, + refine acc.of_downward_closed f' _ _ (_), + { rintro _ ⟨_, c', rfl⟩ -, + exact ⟨c', rfl⟩ }, + { exact h.apply _ } +end + +@[simp] lemma well_founded_on_image {s : set β} : + (f '' s).well_founded_on r ↔ s.well_founded_on (r on f) := +by { rw image_eq_range, exact well_founded_on_range } + namespace well_founded_on protected lemma induction (hs : s.well_founded_on r) (hx : x ∈ s) {P : α → Prop} @@ -98,6 +119,9 @@ begin exact ⟨hle _ _ xy.1, hst xy.2.1, hst xy.2.2⟩ end +lemma mono' (h : ∀ a b ∈ s, r' a b → r a b) : s.well_founded_on r → s.well_founded_on r' := +subrelation.wf $ λ a b, h _ a.2 _ b.2 + lemma subset (h : t.well_founded_on r) (hst : s ⊆ t) : s.well_founded_on r := h.mono le_rfl hst open relation @@ -707,3 +731,71 @@ begin refine ⟨g'.trans g, λ a b hab, (finset.forall_mem_cons _ _).2 _⟩, exact ⟨hg (order_hom_class.mono g' hab), hg' hab⟩ } end + +section prod_lex +variables {rα : α → α → Prop} {rβ : β → β → Prop} {f : γ → α} {g : γ → β} {s : set γ} + +/-- Stronger version of `prod.lex_wf`. Instead of requiring `rβ on g` to be well-founded, we only +require it to be well-founded on fibers of `f`.-/ +lemma well_founded.prod_lex_of_well_founded_on_fiber (hα : well_founded (rα on f)) + (hβ : ∀ a, (f ⁻¹' {a}).well_founded_on (rβ on g)) : + well_founded (prod.lex rα rβ on λ c, (f c, g c)) := +begin + refine ((psigma.lex_wf (well_founded_on_range.2 hα) $ λ a, hβ a).on_fun).mono (λ c c' h, _), + exact λ c, ⟨⟨_, c, rfl⟩, c, rfl⟩, + obtain h' | h' := prod.lex_iff.1 h, + { exact psigma.lex.left _ _ h' }, + { dsimp only [inv_image, (on)] at h' ⊢, + convert psigma.lex.right (⟨_, c', rfl⟩ : range f) _ using 1, swap, + exacts [⟨c, h'.1⟩, psigma.subtype_ext (subtype.ext h'.1) rfl, h'.2] } +end + +lemma set.well_founded_on.prod_lex_of_well_founded_on_fiber (hα : s.well_founded_on (rα on f)) + (hβ : ∀ a, (s ∩ f ⁻¹' {a}).well_founded_on (rβ on g)) : + s.well_founded_on (prod.lex rα rβ on λ c, (f c, g c)) := +begin + refine well_founded.prod_lex_of_well_founded_on_fiber hα + (λ a, subrelation.wf (λ b c h, _) (hβ a).on_fun), + exact λ x, ⟨x, x.1.2, x.2⟩, + assumption, +end + +end prod_lex + +section sigma_lex +variables {rι : ι → ι → Prop} {rπ : Π i, π i → π i → Prop} {f : γ → ι} {g : Π i, γ → π i} + {s : set γ} + +/-- Stronger version of `psigma.lex_wf`. Instead of requiring `rπ on g` to be well-founded, we only +require it to be well-founded on fibers of `f`.-/ +lemma well_founded.sigma_lex_of_well_founded_on_fiber (hι : well_founded (rι on f)) + (hπ : ∀ i, (f ⁻¹' {i}).well_founded_on (rπ i on g i)) : + well_founded (sigma.lex rι rπ on λ c, ⟨f c, g (f c) c⟩) := +begin + refine ((psigma.lex_wf (well_founded_on_range.2 hι) $ λ a, hπ a).on_fun).mono (λ c c' h, _), + exact λ c, ⟨⟨_, c, rfl⟩, c, rfl⟩, + obtain h' | ⟨h', h''⟩ := sigma.lex_iff.1 h, + { exact psigma.lex.left _ _ h' }, + { dsimp only [inv_image, (on)] at h' ⊢, + convert psigma.lex.right (⟨_, c', rfl⟩ : range f) _ using 1, swap, + { exact ⟨c, h'⟩ }, + { exact psigma.subtype_ext (subtype.ext h') rfl }, + { dsimp only [subtype.coe_mk] at *, + revert h', + generalize : f c = d, + rintro rfl _, + exact h'' } } +end + +lemma set.well_founded_on.sigma_lex_of_well_founded_on_fiber (hι : s.well_founded_on (rι on f)) + (hπ : ∀ i, (s ∩ f ⁻¹' {i}).well_founded_on (rπ i on g i)) : + s.well_founded_on (sigma.lex rι rπ on λ c, ⟨f c, g (f c) c⟩) := +begin + show well_founded (sigma.lex rι rπ on λ (c : s), ⟨f c, g (f c) c⟩), + refine @well_founded.sigma_lex_of_well_founded_on_fiber _ s _ _ rπ (λ c, f c) (λ i c, g _ c) hι + (λ i, subrelation.wf (λ b c h, _) (hπ i).on_fun), + exact λ x, ⟨x, x.1.2, x.2⟩, + assumption, +end + +end sigma_lex From 1d29de43a5ba4662dd33b5cfeecfc2a27a5a8a29 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 14 Jul 2023 19:58:49 +0000 Subject: [PATCH 1252/1291] feat(data/*/interval): `finset.uIcc` on concrete structures (#18838) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Calculate the size of `finset.uIcc` in `ℕ`, `ℤ`, `fin`, `prod`, `pi`, `multiset`, `finset`... Co-authored-by: Eric Wieser --- src/data/dfinsupp/interval.lean | 13 ++++++-- src/data/dfinsupp/multiset.lean | 52 ++++++++++++++++++++++++++--- src/data/dfinsupp/order.lean | 29 ++++++++++------ src/data/fin/interval.lean | 20 +++++++++++ src/data/finset/locally_finite.lean | 17 +++++++--- src/data/finsupp/interval.lean | 9 +++++ src/data/finsupp/order.lean | 14 ++++++-- src/data/int/interval.lean | 23 +++++++------ src/data/multiset/interval.lean | 34 ++++++++++++------- src/data/nat/interval.lean | 14 ++++++++ src/data/pi/interval.lean | 43 +++++++++++++----------- src/data/pnat/interval.lean | 20 ++++++++--- src/order/locally_finite.lean | 38 ++++++++++----------- 13 files changed, 234 insertions(+), 92 deletions(-) diff --git a/src/data/dfinsupp/interval.lean b/src/data/dfinsupp/interval.lean index e010360860147..990341f78ec32 100644 --- a/src/data/dfinsupp/interval.lean +++ b/src/data/dfinsupp/interval.lean @@ -140,7 +140,7 @@ end end pi -section locally_finite +section partial_order variables [decidable_eq ι] [Π i, decidable_eq (α i)] variables [Π i, partial_order (α i)] [Π i, has_zero (α i)] [Π i, locally_finite_order (α i)] @@ -169,7 +169,16 @@ by rw [card_Ioc_eq_card_Icc_sub_one, card_Icc] lemma card_Ioo : (Ioo f g).card = ∏ i in f.support ∪ g.support, (Icc (f i) (g i)).card - 2 := by rw [card_Ioo_eq_card_Icc_sub_two, card_Icc] -end locally_finite +end partial_order + +section lattice +variables [decidable_eq ι] [Π i, decidable_eq (α i)] [Π i, lattice (α i)] [Π i, has_zero (α i)] + [Π i, locally_finite_order (α i)] (f g : Π₀ i, α i) + +lemma card_uIcc : (uIcc f g).card = ∏ i in f.support ∪ g.support, (uIcc (f i) (g i)).card := +by { rw ←support_inf_union_support_sup, exact card_Icc _ _ } + +end lattice section canonically_ordered variables [decidable_eq ι] [Π i, decidable_eq (α i)] diff --git a/src/data/dfinsupp/multiset.lean b/src/data/dfinsupp/multiset.lean index de23e647e818c..de2ed3c8f1031 100644 --- a/src/data/dfinsupp/multiset.lean +++ b/src/data/dfinsupp/multiset.lean @@ -17,6 +17,8 @@ with `multiset.to_dfinsupp` the reverse equivalence. Note that this provides a computable alternative to `finsupp.to_multiset`. -/ +open function + variables {α : Type*} {β : α → Type*} namespace dfinsupp @@ -38,7 +40,7 @@ dfinsupp.sum_add_hom_single _ _ _ end dfinsupp namespace multiset -variables [decidable_eq α] +variables [decidable_eq α] {s t : multiset α} /-- A computable version of `multiset.to_finsupp` -/ def to_dfinsupp : multiset α →+ Π₀ a : α, ℕ := @@ -78,12 +80,52 @@ add_monoid_hom.to_add_equiv @[simp] lemma to_dfinsupp_to_multiset (s : multiset α) : s.to_dfinsupp.to_multiset = s := equiv_dfinsupp.symm_apply_apply s -@[simp] lemma to_dfinsupp_le_to_dfinsupp (s t : multiset α) : - to_dfinsupp s ≤ to_dfinsupp t ↔ s ≤ t := +lemma to_dfinsupp_injective : injective (to_dfinsupp : multiset α → Π₀ a, ℕ) := +equiv_dfinsupp.injective + +@[simp] lemma to_dfinsupp_inj : to_dfinsupp s = to_dfinsupp t ↔ s = t := +to_dfinsupp_injective.eq_iff + +@[simp] lemma to_dfinsupp_le_to_dfinsupp : to_dfinsupp s ≤ to_dfinsupp t ↔ s ≤ t := by simp [multiset.le_iff_count, dfinsupp.le_def] +@[simp] lemma to_dfinsupp_lt_to_dfinsupp : to_dfinsupp s < to_dfinsupp t ↔ s < t := +lt_iff_lt_of_le_iff_le' to_dfinsupp_le_to_dfinsupp to_dfinsupp_le_to_dfinsupp + +@[simp] lemma to_dfinsupp_inter (s t : multiset α) : + to_dfinsupp (s ∩ t) = s.to_dfinsupp ⊓ t.to_dfinsupp := +by { ext i, simp [inf_eq_min] } + +@[simp] lemma to_dfinsupp_union (s t : multiset α) : + to_dfinsupp (s ∪ t) = s.to_dfinsupp ⊔ t.to_dfinsupp := +by { ext i, simp [sup_eq_max] } + end multiset -@[simp] lemma dfinsupp.to_multiset_to_dfinsupp [decidable_eq α] (f : Π₀ a : α, ℕ) : - f.to_multiset.to_dfinsupp = f := +namespace dfinsupp +variables [decidable_eq α] {f g : Π₀ a : α, ℕ} + +@[simp] lemma to_multiset_to_dfinsupp : f.to_multiset.to_dfinsupp = f := multiset.equiv_dfinsupp.apply_symm_apply f + +lemma to_multiset_injective : injective (to_multiset : (Π₀ a, ℕ) → multiset α) := +multiset.equiv_dfinsupp.symm.injective + +@[simp] lemma to_multiset_inj : to_multiset f = to_multiset g ↔ f = g := +to_multiset_injective.eq_iff + +@[simp] lemma to_multiset_le_to_multiset : to_multiset f ≤ to_multiset g ↔ f ≤ g := +by simp_rw [←multiset.to_dfinsupp_le_to_dfinsupp, to_multiset_to_dfinsupp] + +@[simp] lemma to_multiset_lt_to_multiset : to_multiset f < to_multiset g ↔ f < g := +by simp_rw [←multiset.to_dfinsupp_lt_to_dfinsupp, to_multiset_to_dfinsupp] + +variables (f g) + +@[simp] lemma to_multiset_inf : to_multiset (f ⊓ g) = f.to_multiset ∩ g.to_multiset := +multiset.to_dfinsupp_injective $ by simp + +@[simp] lemma to_multiset_sup : to_multiset (f ⊔ g) = f.to_multiset ∪ g.to_multiset := +multiset.to_dfinsupp_injective $ by simp + +end dfinsupp diff --git a/src/data/dfinsupp/order.lean b/src/data/dfinsupp/order.lean index 0c6344b670dda..c5a90bc29224c 100644 --- a/src/data/dfinsupp/order.lean +++ b/src/data/dfinsupp/order.lean @@ -31,15 +31,13 @@ namespace dfinsupp /-! ### Order structures -/ section has_zero -variables (α) [Π i, has_zero (α i)] +variables [Π i, has_zero (α i)] section has_le variables [Π i, has_le (α i)] instance : has_le (Π₀ i, α i) := ⟨λ f g, ∀ i, f i ≤ g i⟩ -variables {α} - lemma le_def {f g : Π₀ i, α i} : f ≤ g ↔ ∀ i, f i ≤ g i := iff.rfl /-- The order on `dfinsupp`s over a partial order embeds into the order on functions -/ @@ -59,7 +57,7 @@ variables [Π i, preorder (α i)] instance : preorder (Π₀ i, α i) := { le_refl := λ f i, le_rfl, le_trans := λ f g h hfg hgh i, (hfg i).trans (hgh i), - .. dfinsupp.has_le α } + .. dfinsupp.has_le } lemma coe_fn_mono : monotone (coe_fn : (Π₀ i, α i) → Π i, α i) := λ f g, le_def.1 @@ -67,14 +65,14 @@ end preorder instance [Π i, partial_order (α i)] : partial_order (Π₀ i, α i) := { le_antisymm := λ f g hfg hgf, ext $ λ i, (hfg i).antisymm (hgf i), - .. dfinsupp.preorder α} + .. dfinsupp.preorder } instance [Π i, semilattice_inf (α i)] : semilattice_inf (Π₀ i, α i) := { inf := zip_with (λ _, (⊓)) (λ _, inf_idem), inf_le_left := λ f g i, by { rw zip_with_apply, exact inf_le_left }, inf_le_right := λ f g i, by { rw zip_with_apply, exact inf_le_right }, le_inf := λ f g h hf hg i, by { rw zip_with_apply, exact le_inf (hf i) (hg i) }, - ..dfinsupp.partial_order α } + ..dfinsupp.partial_order } @[simp] lemma inf_apply [Π i, semilattice_inf (α i)] (f g : Π₀ i, α i) (i : ι) : (f ⊓ g) i = f i ⊓ g i := @@ -85,15 +83,26 @@ instance [Π i, semilattice_sup (α i)] : semilattice_sup (Π₀ i, α i) := le_sup_left := λ f g i, by { rw zip_with_apply, exact le_sup_left }, le_sup_right := λ f g i, by { rw zip_with_apply, exact le_sup_right }, sup_le := λ f g h hf hg i, by { rw zip_with_apply, exact sup_le (hf i) (hg i) }, - ..dfinsupp.partial_order α } + ..dfinsupp.partial_order } @[simp] lemma sup_apply [Π i, semilattice_sup (α i)] (f g : Π₀ i, α i) (i : ι) : (f ⊔ g) i = f i ⊔ g i := zip_with_apply _ _ _ _ _ -instance lattice [Π i, lattice (α i)] : lattice (Π₀ i, α i) := -{ .. dfinsupp.semilattice_inf α, .. dfinsupp.semilattice_sup α } +section lattice +variables [Π i, lattice (α i)] (f g : Π₀ i, α i) + +instance lattice : lattice (Π₀ i, α i) := { ..dfinsupp.semilattice_inf, ..dfinsupp.semilattice_sup } + +variables [decidable_eq ι] [Π i (x : α i), decidable (x ≠ 0)] + +lemma support_inf_union_support_sup : (f ⊓ g).support ∪ (f ⊔ g).support = f.support ∪ g.support := +coe_injective $ compl_injective $ by { ext, simp [inf_eq_and_sup_eq_iff] } + +lemma support_sup_union_support_inf : (f ⊔ g).support ∪ (f ⊓ g).support = f.support ∪ g.support := +(union_comm _ _).trans $ support_inf_union_support_sup _ _ +end lattice end has_zero /-! ### Algebraic order structures -/ @@ -102,7 +111,7 @@ instance (α : ι → Type*) [Π i, ordered_add_comm_monoid (α i)] : ordered_add_comm_monoid (Π₀ i, α i) := { add_le_add_left := λ a b h c i, by { rw [add_apply, add_apply], exact add_le_add_left (h i) (c i) }, - .. dfinsupp.add_comm_monoid, .. dfinsupp.partial_order α } + .. dfinsupp.add_comm_monoid, .. dfinsupp.partial_order } instance (α : ι → Type*) [Π i, ordered_cancel_add_comm_monoid (α i)] : ordered_cancel_add_comm_monoid (Π₀ i, α i) := diff --git a/src/data/fin/interval.lean b/src/data/fin/interval.lean index 41a0f2409e519..46e3a23c11bdf 100644 --- a/src/data/fin/interval.lean +++ b/src/data/fin/interval.lean @@ -16,6 +16,16 @@ This file proves that `fin n` is a `locally_finite_order` and calculates the car intervals as finsets and fintypes. -/ +namespace fin +variables {n : ℕ} (a b : fin n) + +@[simp, norm_cast] lemma coe_sup : (↑(a ⊔ b) : ℕ) = a ⊔ b := rfl +@[simp, norm_cast] lemma coe_inf : (↑(a ⊓ b) : ℕ) = a ⊓ b := rfl +@[simp, norm_cast] lemma coe_max : (↑(max a b) : ℕ) = max a b := rfl +@[simp, norm_cast] lemma coe_min : (↑(min a b) : ℕ) = min a b := rfl + +end fin + open finset fin function open_locale big_operators @@ -39,6 +49,7 @@ lemma Icc_eq_finset_subtype : Icc a b = (Icc (a : ℕ) b).fin n := rfl lemma Ico_eq_finset_subtype : Ico a b = (Ico (a : ℕ) b).fin n := rfl lemma Ioc_eq_finset_subtype : Ioc a b = (Ioc (a : ℕ) b).fin n := rfl lemma Ioo_eq_finset_subtype : Ioo a b = (Ioo (a : ℕ) b).fin n := rfl +lemma uIcc_eq_finset_subtype : uIcc a b = (uIcc (a : ℕ) b).fin n := rfl @[simp] lemma map_subtype_embedding_Icc : (Icc a b).map fin.coe_embedding = Icc a b := by simp [Icc_eq_finset_subtype, finset.fin, finset.map_map, Icc_filter_lt_of_lt_right] @@ -52,6 +63,9 @@ by simp [Ioc_eq_finset_subtype, finset.fin, finset.map_map, Ioc_filter_lt_of_lt_ @[simp] lemma map_subtype_embedding_Ioo : (Ioo a b).map fin.coe_embedding = Ioo a b := by simp [Ioo_eq_finset_subtype, finset.fin, finset.map_map] +@[simp] lemma map_subtype_embedding_uIcc : (uIcc a b).map coe_embedding = uIcc a b := +map_subtype_embedding_Icc _ _ + @[simp] lemma card_Icc : (Icc a b).card = b + 1 - a := by rw [←nat.card_Icc, ←map_subtype_embedding_Icc, card_map] @@ -64,6 +78,9 @@ by rw [←nat.card_Ioc, ←map_subtype_embedding_Ioc, card_map] @[simp] lemma card_Ioo : (Ioo a b).card = b - a - 1 := by rw [←nat.card_Ioo, ←map_subtype_embedding_Ioo, card_map] +@[simp] lemma card_uIcc : (uIcc a b).card = (b - a : ℤ).nat_abs + 1 := +by rw [coe_coe, coe_coe, ←nat.card_uIcc, ←map_subtype_embedding_uIcc, card_map] + @[simp] lemma card_fintype_Icc : fintype.card (set.Icc a b) = b + 1 - a := by rw [←card_Icc, fintype.card_of_finset] @@ -76,6 +93,9 @@ by rw [←card_Ioc, fintype.card_of_finset] @[simp] lemma card_fintype_Ioo : fintype.card (set.Ioo a b) = b - a - 1 := by rw [←card_Ioo, fintype.card_of_finset] +@[simp] lemma card_fintype_uIcc : fintype.card (set.uIcc a b) = (b - a : ℤ).nat_abs + 1 := +by rw [←card_uIcc, fintype.card_of_finset] + lemma Ici_eq_finset_subtype : Ici a = (Icc (a : ℕ) n).fin n := by { ext, simp } lemma Ioi_eq_finset_subtype : Ioi a = (Ioc (a : ℕ) n).fin n := by { ext, simp } lemma Iic_eq_finset_subtype : Iic b = (Iic (b : ℕ)).fin n := rfl diff --git a/src/data/finset/locally_finite.lean b/src/data/finset/locally_finite.lean index 8dc178e34bc1d..f8b6e5b429cc9 100644 --- a/src/data/finset/locally_finite.lean +++ b/src/data/finset/locally_finite.lean @@ -564,8 +564,11 @@ lemma Icc_subset_uIcc' : Icc b a ⊆ [a, b] := Icc_subset_Icc inf_le_right le_su @[simp] lemma left_mem_uIcc : a ∈ [a, b] := mem_Icc.2 ⟨inf_le_left, le_sup_left⟩ @[simp] lemma right_mem_uIcc : b ∈ [a, b] := mem_Icc.2 ⟨inf_le_right, le_sup_right⟩ -lemma mem_uIcc_of_le (ha : a ≤ x) (hb : x ≤ b) : x ∈ [a, b] := Icc_subset_uIcc $ mem_Icc.2 ⟨ha, hb⟩ -lemma mem_uIcc_of_ge (hb : b ≤ x) (ha : x ≤ a) : x ∈ [a, b] := Icc_subset_uIcc' $ mem_Icc.2 ⟨hb, ha⟩ +lemma mem_uIcc_of_le (ha : a ≤ x) (hb : x ≤ b) : x ∈ [a, b] := +Icc_subset_uIcc $ mem_Icc.2 ⟨ha, hb⟩ + +lemma mem_uIcc_of_ge (hb : b ≤ x) (ha : x ≤ a) : x ∈ [a, b] := +Icc_subset_uIcc' $ mem_Icc.2 ⟨hb, ha⟩ lemma uIcc_subset_uIcc (h₁ : a₁ ∈ [a₂, b₂]) (h₂ : b₁ ∈ [a₂, b₂]) : [a₁, b₁] ⊆ [a₂, b₂] := by { rw mem_uIcc at h₁ h₂, exact Icc_subset_Icc (le_inf h₁.1 h₂.1) (sup_le h₁.2 h₂.2) } @@ -576,11 +579,15 @@ by { rw mem_Icc at ha hb, exact Icc_subset_Icc (le_inf ha.1 hb.1) (sup_le ha.2 h lemma uIcc_subset_uIcc_iff_mem : [a₁, b₁] ⊆ [a₂, b₂] ↔ a₁ ∈ [a₂, b₂] ∧ b₁ ∈ [a₂, b₂] := ⟨λ h, ⟨h left_mem_uIcc, h right_mem_uIcc⟩, λ h, uIcc_subset_uIcc h.1 h.2⟩ -lemma uIcc_subset_uIcc_iff_le' : [a₁, b₁] ⊆ [a₂, b₂] ↔ a₂ ⊓ b₂ ≤ a₁ ⊓ b₁ ∧ a₁ ⊔ b₁ ≤ a₂ ⊔ b₂ := +lemma uIcc_subset_uIcc_iff_le' : + [a₁, b₁] ⊆ [a₂, b₂] ↔ a₂ ⊓ b₂ ≤ a₁ ⊓ b₁ ∧ a₁ ⊔ b₁ ≤ a₂ ⊔ b₂ := Icc_subset_Icc_iff inf_le_sup -lemma uIcc_subset_uIcc_right (h : x ∈ [a, b]) : [x, b] ⊆ [a, b] := uIcc_subset_uIcc h right_mem_uIcc -lemma uIcc_subset_uIcc_left (h : x ∈ [a, b]) : [a, x] ⊆ [a, b] := uIcc_subset_uIcc left_mem_uIcc h +lemma uIcc_subset_uIcc_right (h : x ∈ [a, b]) : [x, b] ⊆ [a, b] := +uIcc_subset_uIcc h right_mem_uIcc + +lemma uIcc_subset_uIcc_left (h : x ∈ [a, b]) : [a, x] ⊆ [a, b] := +uIcc_subset_uIcc left_mem_uIcc h end lattice diff --git a/src/data/finsupp/interval.lean b/src/data/finsupp/interval.lean index b71f5d4d4987c..b05e89043e75b 100644 --- a/src/data/finsupp/interval.lean +++ b/src/data/finsupp/interval.lean @@ -106,6 +106,15 @@ by rw [card_Ioo_eq_card_Icc_sub_two, card_Icc] end partial_order +section lattice +variables [lattice α] [has_zero α] [locally_finite_order α] (f g : ι →₀ α) + +lemma card_uIcc [decidable_eq ι] : + (uIcc f g).card = ∏ i in f.support ∪ g.support, (uIcc (f i) (g i)).card := +by { rw ←support_inf_union_support_sup, exact card_Icc _ _ } + +end lattice + section canonically_ordered variables [canonically_ordered_add_monoid α] [locally_finite_order α] diff --git a/src/data/finsupp/order.lean b/src/data/finsupp/order.lean index 889ee49238192..b5aa762727bc9 100644 --- a/src/data/finsupp/order.lean +++ b/src/data/finsupp/order.lean @@ -86,8 +86,18 @@ instance [semilattice_sup α] : semilattice_sup (ι →₀ α) := @[simp] lemma sup_apply [semilattice_sup α] {i : ι} {f g : ι →₀ α} : (f ⊔ g) i = f i ⊔ g i := rfl -instance lattice [lattice α] : lattice (ι →₀ α) := -{ .. finsupp.semilattice_inf, .. finsupp.semilattice_sup } +instance [lattice α] : lattice (ι →₀ α) := { ..finsupp.semilattice_inf, ..finsupp.semilattice_sup } + +section lattice +variables [decidable_eq ι] [lattice α] (f g : ι →₀ α) + +lemma support_inf_union_support_sup : (f ⊓ g).support ∪ (f ⊔ g).support = f.support ∪ g.support := +coe_injective $ compl_injective $ by { ext, simp [inf_eq_and_sup_eq_iff] } + +lemma support_sup_union_support_inf : (f ⊔ g).support ∪ (f ⊓ g).support = f.support ∪ g.support := +(union_comm _ _).trans $ support_inf_union_support_sup _ _ + +end lattice end has_zero diff --git a/src/data/int/interval.lean b/src/data/int/interval.lean index 4856c9e400009..4b0c1c7b84641 100644 --- a/src/data/int/interval.lean +++ b/src/data/int/interval.lean @@ -92,18 +92,18 @@ lemma Ioc_eq_finset_map : lemma Ioo_eq_finset_map : Ioo a b = (finset.range (b - a - 1).to_nat).map (nat.cast_embedding.trans $ add_left_embedding (a + 1)) := rfl +lemma uIcc_eq_finset_map : uIcc a b = (range (max a b + 1 - min a b).to_nat).map + (nat.cast_embedding.trans $ add_left_embedding $ min a b) := rfl -@[simp] lemma card_Icc : (Icc a b).card = (b + 1 - a).to_nat := -by { change (finset.map _ _).card = _, rw [finset.card_map, finset.card_range] } +@[simp] lemma card_Icc : (Icc a b).card = (b + 1 - a).to_nat := (card_map _).trans $ card_range _ +@[simp] lemma card_Ico : (Ico a b).card = (b - a).to_nat := (card_map _).trans $ card_range _ +@[simp] lemma card_Ioc : (Ioc a b).card = (b - a).to_nat := (card_map _).trans $ card_range _ +@[simp] lemma card_Ioo : (Ioo a b).card = (b - a - 1).to_nat := (card_map _).trans $ card_range _ -@[simp] lemma card_Ico : (Ico a b).card = (b - a).to_nat := -by { change (finset.map _ _).card = _, rw [finset.card_map, finset.card_range] } - -@[simp] lemma card_Ioc : (Ioc a b).card = (b - a).to_nat := -by { change (finset.map _ _).card = _, rw [finset.card_map, finset.card_range] } - -@[simp] lemma card_Ioo : (Ioo a b).card = (b - a - 1).to_nat := -by { change (finset.map _ _).card = _, rw [finset.card_map, finset.card_range] } +@[simp] lemma card_uIcc : (uIcc a b).card = (b - a).nat_abs + 1 := +(card_map _).trans $ int.coe_nat_inj $ by rw [card_range, sup_eq_max, inf_eq_min, + int.to_nat_of_nonneg (sub_nonneg_of_le $ le_add_one min_le_max), int.coe_nat_add, int.coe_nat_abs, + add_comm, add_sub_assoc, max_sub_min_eq_abs, add_comm, int.coe_nat_one] lemma card_Icc_of_le (h : a ≤ b + 1) : ((Icc a b).card : ℤ) = b + 1 - a := by rw [card_Icc, to_nat_sub_of_le h] @@ -129,6 +129,9 @@ by rw [←card_Ioc, fintype.card_of_finset] @[simp] lemma card_fintype_Ioo : fintype.card (set.Ioo a b) = (b - a - 1).to_nat := by rw [←card_Ioo, fintype.card_of_finset] +@[simp] lemma card_fintype_uIcc : fintype.card (set.uIcc a b) = (b - a).nat_abs + 1 := +by rw [←card_uIcc, fintype.card_of_finset] + lemma card_fintype_Icc_of_le (h : a ≤ b + 1) : (fintype.card (set.Icc a b) : ℤ) = b + 1 - a := by rw [card_fintype_Icc, to_nat_sub_of_le h] diff --git a/src/data/multiset/interval.lean b/src/data/multiset/interval.lean index 33ae1fb363183..8e36ee02b2efb 100644 --- a/src/data/multiset/interval.lean +++ b/src/data/multiset/interval.lean @@ -29,42 +29,52 @@ multisets are typically used computationally. open finset dfinsupp function open_locale big_operators pointwise -variables {α : Type*} {β : α → Type*} +variables {α : Type*} namespace multiset - -variables [decidable_eq α] (f g : multiset α) +variables [decidable_eq α] (s t : multiset α) instance : locally_finite_order (multiset α) := locally_finite_order.of_Icc (multiset α) - (λ f g, (finset.Icc f.to_dfinsupp g.to_dfinsupp).map + (λ s t, (finset.Icc s.to_dfinsupp t.to_dfinsupp).map (multiset.equiv_dfinsupp.to_equiv.symm.to_embedding)) - (λ f g x, by simp) + (λ s t x, by simp) lemma Icc_eq : - finset.Icc f g = - (finset.Icc f.to_dfinsupp g.to_dfinsupp).map + finset.Icc s t = + (finset.Icc s.to_dfinsupp t.to_dfinsupp).map (multiset.equiv_dfinsupp.to_equiv.symm.to_embedding) := rfl +lemma uIcc_eq : + uIcc s t = + (uIcc s.to_dfinsupp t.to_dfinsupp).map + (multiset.equiv_dfinsupp.to_equiv.symm.to_embedding) := +(Icc_eq _ _).trans $ by simp [uIcc] + lemma card_Icc : - (finset.Icc f g).card = ∏ i in f.to_finset ∪ g.to_finset, (g.count i + 1 - f.count i) := + (finset.Icc s t).card = ∏ i in s.to_finset ∪ t.to_finset, (t.count i + 1 - s.count i) := by simp_rw [Icc_eq, finset.card_map, dfinsupp.card_Icc, nat.card_Icc, multiset.to_dfinsupp_apply, to_dfinsupp_support] lemma card_Ico : - (finset.Ico f g).card = ∏ i in f.to_finset ∪ g.to_finset, (g.count i + 1 - f.count i) - 1 := + (finset.Ico s t).card = ∏ i in s.to_finset ∪ t.to_finset, (t.count i + 1 - s.count i) - 1 := by rw [card_Ico_eq_card_Icc_sub_one, card_Icc] lemma card_Ioc : - (finset.Ioc f g).card = ∏ i in f.to_finset ∪ g.to_finset, (g.count i + 1 - f.count i) - 1 := + (finset.Ioc s t).card = ∏ i in s.to_finset ∪ t.to_finset, (t.count i + 1 - s.count i) - 1 := by rw [card_Ioc_eq_card_Icc_sub_one, card_Icc] lemma card_Ioo : - (finset.Ioo f g).card = ∏ i in f.to_finset ∪ g.to_finset, (g.count i + 1 - f.count i) - 2 := + (finset.Ioo s t).card = ∏ i in s.to_finset ∪ t.to_finset, (t.count i + 1 - s.count i) - 2 := by rw [card_Ioo_eq_card_Icc_sub_two, card_Icc] +lemma card_uIcc : + (uIcc s t).card = ∏ i in s.to_finset ∪ t.to_finset, ((t.count i - s.count i : ℤ).nat_abs + 1) := +by simp_rw [uIcc_eq, finset.card_map, dfinsupp.card_uIcc, nat.card_uIcc, multiset.to_dfinsupp_apply, + to_dfinsupp_support] + lemma card_Iic : - (finset.Iic f).card = ∏ i in f.to_finset, (f.count i + 1) := + (finset.Iic s).card = ∏ i in s.to_finset, (s.count i + 1) := by simp_rw [Iic_eq_Icc, card_Icc, bot_eq_zero, to_finset_zero, empty_union, count_zero, tsub_zero] end multiset diff --git a/src/data/nat/interval.lean b/src/data/nat/interval.lean index d1a5a2e862753..a7efc2bf8d370 100644 --- a/src/data/nat/interval.lean +++ b/src/data/nat/interval.lean @@ -65,6 +65,8 @@ lemma Icc_eq_range' : Icc a b = ⟨list.range' a (b + 1 - a), list.nodup_range' lemma Ico_eq_range' : Ico a b = ⟨list.range' a (b - a), list.nodup_range' _ _⟩ := rfl lemma Ioc_eq_range' : Ioc a b = ⟨list.range' (a + 1) (b - a), list.nodup_range' _ _⟩ := rfl lemma Ioo_eq_range' : Ioo a b = ⟨list.range' (a + 1) (b - a - 1), list.nodup_range' _ _⟩ := rfl +lemma uIcc_eq_range' : + uIcc a b = ⟨list.range' (min a b) (max a b + 1 - min a b), list.nodup_range' _ _⟩ := rfl lemma Iio_eq_range : Iio = range := by { ext b x, rw [mem_Iio, mem_range] } @@ -76,6 +78,18 @@ lemma _root_.finset.range_eq_Ico : range = Ico 0 := Ico_zero_eq_range.symm @[simp] lemma card_Ico : (Ico a b).card = b - a := list.length_range' _ _ @[simp] lemma card_Ioc : (Ioc a b).card = b - a := list.length_range' _ _ @[simp] lemma card_Ioo : (Ioo a b).card = b - a - 1 := list.length_range' _ _ + +@[simp] lemma card_uIcc : (uIcc a b).card = (b - a : ℤ).nat_abs + 1 := +begin + refine (card_Icc _ _).trans (int.coe_nat_inj _), + rw [sup_eq_max, inf_eq_min, int.coe_nat_sub], + { rw [add_comm, int.coe_nat_add, add_sub_assoc], + norm_cast, + push_cast, + rw [max_sub_min_eq_abs, add_comm] }, + { exact min_le_max.trans le_self_add } +end + @[simp] lemma card_Iic : (Iic b).card = b + 1 := by rw [Iic_eq_Icc, card_Icc, bot_eq_zero, tsub_zero] diff --git a/src/data/pi/interval.lean b/src/data/pi/interval.lean index c77e21f3ce2bd..812b52ef97122 100644 --- a/src/data/pi/interval.lean +++ b/src/data/pi/interval.lean @@ -19,14 +19,14 @@ order are locally finite and calculates the cardinality of their intervals. open finset fintype open_locale big_operators -variables {ι : Type*} {α : ι → Type*} - +variables {ι : Type*} {α : ι → Type*} [fintype ι] [decidable_eq ι] [Π i, decidable_eq (α i)] namespace pi +section partial_order +variables [Π i, partial_order (α i)] -section locally_finite -variables [decidable_eq ι] [fintype ι] [Π i, decidable_eq (α i)] - [Π i, partial_order (α i)] [Π i, locally_finite_order (α i)] +section locally_finite_order +variables [Π i, locally_finite_order (α i)] instance : locally_finite_order (Π i, α i) := locally_finite_order.of_Icc _ @@ -39,21 +39,18 @@ lemma Icc_eq : Icc a b = pi_finset (λ i, Icc (a i) (b i)) := rfl lemma card_Icc : (Icc a b).card = ∏ i, (Icc (a i) (b i)).card := card_pi_finset _ -lemma card_Ico : (Ico a b).card = (∏ i, (Icc (a i) (b i)).card) - 1 := +lemma card_Ico : (Ico a b).card = ∏ i, (Icc (a i) (b i)).card - 1 := by rw [card_Ico_eq_card_Icc_sub_one, card_Icc] -lemma card_Ioc : (Ioc a b).card = (∏ i, (Icc (a i) (b i)).card) - 1 := +lemma card_Ioc : (Ioc a b).card = ∏ i, (Icc (a i) (b i)).card - 1 := by rw [card_Ioc_eq_card_Icc_sub_one, card_Icc] -lemma card_Ioo : (Ioo a b).card = (∏ i, (Icc (a i) (b i)).card) - 2 := +lemma card_Ioo : (Ioo a b).card = ∏ i, (Icc (a i) (b i)).card - 2 := by rw [card_Ioo_eq_card_Icc_sub_two, card_Icc] -end locally_finite - -section bounded -variables [decidable_eq ι] [fintype ι] [Π i, decidable_eq (α i)] [Π i, partial_order (α i)] +end locally_finite_order -section bot +section locally_finite_order_bot variables [Π i, locally_finite_order_bot (α i)] (b : Π i, α i) instance : locally_finite_order_bot (Π i, α i) := @@ -63,12 +60,12 @@ locally_finite_order_top.of_Iic _ lemma card_Iic : (Iic b).card = ∏ i, (Iic (b i)).card := card_pi_finset _ -lemma card_Iio : (Iio b).card = (∏ i, (Iic (b i)).card) - 1 := +lemma card_Iio : (Iio b).card = ∏ i, (Iic (b i)).card - 1 := by rw [card_Iio_eq_card_Iic_sub_one, card_Iic] -end bot +end locally_finite_order_bot -section top +section locally_finite_order_top variables [Π i, locally_finite_order_top (α i)] (a : Π i, α i) instance : locally_finite_order_top (Π i, α i) := @@ -76,13 +73,19 @@ locally_finite_order_top.of_Ici _ (λ a, pi_finset $ λ i, Ici (a i)) (λ a x, by simp_rw [mem_pi_finset, mem_Ici, le_def]) -lemma card_Ici : (Ici a).card = (∏ i, (Ici (a i)).card) := card_pi_finset _ +lemma card_Ici : (Ici a).card = ∏ i, (Ici (a i)).card := card_pi_finset _ -lemma card_Ioi : (Ioi a).card = (∏ i, (Ici (a i)).card) - 1 := +lemma card_Ioi : (Ioi a).card = ∏ i, (Ici (a i)).card - 1 := by rw [card_Ioi_eq_card_Ici_sub_one, card_Ici] -end top +end locally_finite_order_top +end partial_order + +section lattice +variables [Π i, lattice (α i)] [Π i, locally_finite_order (α i)] (a b : Π i, α i) -end bounded +lemma uIcc_eq : uIcc a b = pi_finset (λ i, uIcc (a i) (b i)) := rfl +lemma card_uIcc : (uIcc a b).card = ∏ i, (uIcc (a i) (b i)).card := card_Icc _ _ +end lattice end pi diff --git a/src/data/pnat/interval.lean b/src/data/pnat/interval.lean index 6d2ecf63145a6..4b6737613cc9f 100644 --- a/src/data/pnat/interval.lean +++ b/src/data/pnat/interval.lean @@ -16,7 +16,7 @@ This file proves that `ℕ+` is a `locally_finite_order` and calculates the card intervals as finsets and fintypes. -/ -open finset pnat +open finset function pnat instance : locally_finite_order ℕ+ := subtype.locally_finite_order _ @@ -27,19 +27,23 @@ lemma Icc_eq_finset_subtype : Icc a b = (Icc (a : ℕ) b).subtype (λ (n : ℕ), lemma Ico_eq_finset_subtype : Ico a b = (Ico (a : ℕ) b).subtype (λ (n : ℕ), 0 < n) := rfl lemma Ioc_eq_finset_subtype : Ioc a b = (Ioc (a : ℕ) b).subtype (λ (n : ℕ), 0 < n) := rfl lemma Ioo_eq_finset_subtype : Ioo a b = (Ioo (a : ℕ) b).subtype (λ (n : ℕ), 0 < n) := rfl +lemma uIcc_eq_finset_subtype : uIcc a b = (uIcc (a : ℕ) b).subtype (λ (n : ℕ), 0 < n) := rfl -lemma map_subtype_embedding_Icc : (Icc a b).map (function.embedding.subtype _) = Icc (a : ℕ) b := +lemma map_subtype_embedding_Icc : (Icc a b).map (embedding.subtype _) = Icc a b := map_subtype_embedding_Icc _ _ _ (λ c _ x hx _ hc _, hc.trans_le hx) -lemma map_subtype_embedding_Ico : (Ico a b).map (function.embedding.subtype _) = Ico (a : ℕ) b := +lemma map_subtype_embedding_Ico : (Ico a b).map (embedding.subtype _) = Ico a b := map_subtype_embedding_Ico _ _ _ (λ c _ x hx _ hc _, hc.trans_le hx) -lemma map_subtype_embedding_Ioc : (Ioc a b).map (function.embedding.subtype _) = Ioc (a : ℕ) b := +lemma map_subtype_embedding_Ioc : (Ioc a b).map (embedding.subtype _) = Ioc a b := map_subtype_embedding_Ioc _ _ _ (λ c _ x hx _ hc _, hc.trans_le hx) -lemma map_subtype_embedding_Ioo : (Ioo a b).map (function.embedding.subtype _) = Ioo (a : ℕ) b := +lemma map_subtype_embedding_Ioo : (Ioo a b).map (embedding.subtype _) = Ioo a b := map_subtype_embedding_Ioo _ _ _ (λ c _ x hx _ hc _, hc.trans_le hx) +lemma map_subtype_embedding_uIcc : (uIcc a b).map (embedding.subtype _) = uIcc a b := +map_subtype_embedding_Icc _ _ + @[simp] lemma card_Icc : (Icc a b).card = b + 1 - a := by rw [←nat.card_Icc, ←map_subtype_embedding_Icc, card_map] @@ -52,6 +56,9 @@ by rw [←nat.card_Ioc, ←map_subtype_embedding_Ioc, card_map] @[simp] lemma card_Ioo : (Ioo a b).card = b - a - 1 := by rw [←nat.card_Ioo, ←map_subtype_embedding_Ioo, card_map] +@[simp] lemma card_uIcc : (uIcc a b).card = (b - a : ℤ).nat_abs + 1 := +by rw [coe_coe, coe_coe, ←nat.card_uIcc, ←map_subtype_embedding_uIcc, card_map] + @[simp] lemma card_fintype_Icc : fintype.card (set.Icc a b) = b + 1 - a := by rw [←card_Icc, fintype.card_of_finset] @@ -64,4 +71,7 @@ by rw [←card_Ioc, fintype.card_of_finset] @[simp] lemma card_fintype_Ioo : fintype.card (set.Ioo a b) = b - a - 1 := by rw [←card_Ioo, fintype.card_of_finset] +@[simp] lemma card_fintype_uIcc : fintype.card (set.uIcc a b) = (b - a : ℤ).nat_abs + 1 := +by rw [←card_uIcc, fintype.card_of_finset] + end pnat diff --git a/src/order/locally_finite.lean b/src/order/locally_finite.lean index 40c3ec825ac84..f6194e68b8b81 100644 --- a/src/order/locally_finite.lean +++ b/src/order/locally_finite.lean @@ -458,17 +458,10 @@ namespace set section preorder variables [preorder α] [locally_finite_order α] (a b : α) -instance fintype_Icc : fintype (Icc a b) := -fintype.of_finset (finset.Icc a b) (λ x, by rw [finset.mem_Icc, mem_Icc]) - -instance fintype_Ico : fintype (Ico a b) := -fintype.of_finset (finset.Ico a b) (λ x, by rw [finset.mem_Ico, mem_Ico]) - -instance fintype_Ioc : fintype (Ioc a b) := -fintype.of_finset (finset.Ioc a b) (λ x, by rw [finset.mem_Ioc, mem_Ioc]) - -instance fintype_Ioo : fintype (Ioo a b) := -fintype.of_finset (finset.Ioo a b) (λ x, by rw [finset.mem_Ioo, mem_Ioo]) +instance fintype_Icc : fintype (Icc a b) := fintype.of_finset (finset.Icc a b) $ λ x, finset.mem_Icc +instance fintype_Ico : fintype (Ico a b) := fintype.of_finset (finset.Ico a b) $ λ x, finset.mem_Ico +instance fintype_Ioc : fintype (Ioc a b) := fintype.of_finset (finset.Ioc a b) $ λ x, finset.mem_Ioc +instance fintype_Ioo : fintype (Ioo a b) := fintype.of_finset (finset.Ioo a b) $ λ x, finset.mem_Ioo lemma finite_Icc : (Icc a b).finite := (Icc a b).to_finite lemma finite_Ico : (Ico a b).finite := (Ico a b).to_finite @@ -480,11 +473,8 @@ end preorder section order_top variables [preorder α] [locally_finite_order_top α] (a : α) -instance fintype_Ici : fintype (Ici a) := -fintype.of_finset (finset.Ici a) (λ x, by rw [finset.mem_Ici, mem_Ici]) - -instance fintype_Ioi : fintype (Ioi a) := -fintype.of_finset (finset.Ioi a) (λ x, by rw [finset.mem_Ioi, mem_Ioi]) +instance fintype_Ici : fintype (Ici a) := fintype.of_finset (finset.Ici a) $ λ x, finset.mem_Ici +instance fintype_Ioi : fintype (Ioi a) := fintype.of_finset (finset.Ioi a) $ λ x, finset.mem_Ioi lemma finite_Ici : (Ici a).finite := (Ici a).to_finite lemma finite_Ioi : (Ioi a).finite := (Ioi a).to_finite @@ -494,17 +484,23 @@ end order_top section order_bot variables [preorder α] [locally_finite_order_bot α] (b : α) -instance fintype_Iic : fintype (Iic b) := -fintype.of_finset (finset.Iic b) (λ x, by rw [finset.mem_Iic, mem_Iic]) - -instance fintype_Iio : fintype (Iio b) := -fintype.of_finset (finset.Iio b) (λ x, by rw [finset.mem_Iio, mem_Iio]) +instance fintype_Iic : fintype (Iic b) := fintype.of_finset (finset.Iic b) $ λ x, finset.mem_Iic +instance fintype_Iio : fintype (Iio b) := fintype.of_finset (finset.Iio b) $ λ x, finset.mem_Iio lemma finite_Iic : (Iic b).finite := (Iic b).to_finite lemma finite_Iio : (Iio b).finite := (Iio b).to_finite end order_bot +section lattice +variables [lattice α] [locally_finite_order α] (a b : α) + +instance fintype_uIcc : fintype (uIcc a b) := +fintype.of_finset (finset.uIcc a b) $ λ x, finset.mem_uIcc + +@[simp] lemma finite_interval : (uIcc a b).finite := (uIcc _ _).to_finite + +end lattice end set /-! ### Instances -/ From 6b711d2ba5d470c040677ddda0c26b0d72283886 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sat, 15 Jul 2023 00:02:14 +0000 Subject: [PATCH 1253/1291] feat(data/matrix/auto): lemmas for arbitrary concrete matrices generated via auto_params (#15738) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This adds a lemma `matrix.mul_fin` that works for arbitrary dimensions of concrete matrices indexed by `fin`. It uses a similar strategy to `category_theory.reassoc_of` to have a tactic populate wildcards in the lemma statement. For example: ```lean example {α} [add_comm_monoid α] [has_mul α] (a₁₁ a₁₂ a₂₁ a₂₂ b₁₁ b₁₂ b₂₁ b₂₂ : α) : !![a₁₁, a₁₂; a₂₁, a₂₂] ⬝ !![b₁₁, b₁₂; b₂₁, b₂₂] = !![a₁₁ * b₁₁ + a₁₂ * b₂₁, a₁₁ * b₁₂ + a₁₂ * b₂₂; a₂₁ * b₁₁ + a₂₂ * b₂₁, a₂₁ * b₁₂ + a₂₂ * b₂₂] := begin rw mul_fin, end example {α} [add_comm_monoid α] [has_mul α] (a₁₁ a₁₂ b₁₁ b₁₂ b₂₁ b₂₂ : α) : !![a₁₁, a₁₂] ⬝ !![b₁₁, b₁₂; b₂₁, b₂₂] = !![a₁₁ * b₁₁ + a₁₂ * b₂₁, a₁₁ * b₁₂ + a₁₂ * b₂₂;] := begin rw mul_fin, end ``` Relevant zulip threads: * [Explicit vector unfolding](https://leanprover.zulipchat.com/#narrow/stream/113489-new-members/topic/Explicit.20vector.20unfolding) * [concrete matrix multiplication](https://leanprover.zulipchat.com/#narrow/stream/217875-Is-there-code-for-X.3F/topic/concrete.20matrix.20multiplication/near/291208624) * [A tactic for expanding matrices into coefficients](https://leanprover.zulipchat.com/#narrow/stream/239415-metaprogramming-.2F-tactics/topic/A.20tactic.20for.20expanding.20matrices.20into.20coefficients) --- scripts/lint-style.py | 2 +- src/algebra/expr.lean | 48 +++++++++ src/data/matrix/auto.lean | 196 ++++++++++++++++++++++++++++++++++++ test/matrix_reflection.lean | 21 ++++ 4 files changed, 266 insertions(+), 1 deletion(-) create mode 100644 src/algebra/expr.lean create mode 100644 src/data/matrix/auto.lean diff --git a/scripts/lint-style.py b/scripts/lint-style.py index 6fbd777019874..c644ebe292a35 100755 --- a/scripts/lint-style.py +++ b/scripts/lint-style.py @@ -133,7 +133,7 @@ def small_alpha_vrachy_check(lines, path): def instance_check(lines, path): errors = [] for line_nr, line in skip_string(skip_comments(enumerate(lines, 1))): - if re.search(r'\b(_inst_)\d+', line) is not None: + if re.search(r'\b(? vector.mmap f ⟨list.fin_range n, list.length_fin_range _⟩ + +namespace matrix + +section fin_eta + +/-- Prove a statement of the form `∀ A : matrix m n α, A = !![A 0 0, ...]`. +Returns the type of this statement and its proof. -/ +meta def fin_eta.prove (m n : ℕ) : tactic (expr × expr) := +do + u ← tactic.mk_meta_univ, + α ← tactic.mk_local' `α binder_info.implicit (expr.sort u.succ), + A ← tactic.mk_local' `A binder_info.default + (expr.const `matrix [level.zero, level.zero, u] `(fin %%`(m)) `(fin %%`(n)) α), + let entries := λ (i : fin m) (j : fin n), A `(i) `(j), + let entry_vals := pi_fin.to_pexpr (λ i, pi_fin.to_pexpr (λ j, to_pexpr $ entries i j)), + let A_eta := (``(@matrix.of (fin %%`(m)) (fin %%`(n)) _).app entry_vals), + A_eq ← tactic.to_expr ``(%%A = %%A_eta), + t ← tactic.pis [α, A] A_eq, + ((), pr) ← tactic.solve_aux t `[intros α A, exact (matrix.eta_expand_eq A).symm], + pure (t, pr) + +/-- Helper tactic used as an `auto_param` for `matrix.fin_eta` -/ +meta def fin_eta.derive : tactic unit := +do + target@`(%%A' = %%A_eta) ← tactic.target, + (expr.const `matrix ls, [`(fin %%m), `(fin %%n), α]) + ← expr.get_app_fn_args <$> tactic.infer_type A', + some (m, n) ← pure (prod.mk <$> m.to_nat <*> n.to_nat) | + fail!"Dimensions {m} {n} are not numerals", + (t,pr) ← matrix.fin_eta.prove m n, + + tactic.unify target (t.instantiate_pis [α, A']), + tactic.exact (pr α A') + +/-- This lemma expands `A` into `!![A 0 0, ...]`. -/ +theorem fin_eta {α} {m n : ℕ} + (A : matrix (fin m) (fin n) α) {«!![A 0 0, ...]» : matrix (fin m) (fin n) α} + (h : A = «!![A 0 0, ...]» . matrix.fin_eta.derive) : A = «!![A 0 0, ...]» := h + +example : true := +begin + let B : matrix (fin 20) (fin 20) ℕ := 0, + have := matrix.fin_eta B, -- 400 coefficients, but very fast + have : B = B, by rw this, + trivial, +end + +end fin_eta + +section of_mul_of_fin + +/-- Choose a name suffix for a matrix index -/ +private def name_suffix {m n : ℕ} : fin m → fin n → string := +let chars := "₀₁₂₃₄₅₆₇₈₉".data in +if h : m ≤ 10 ∧ n ≤ 10 +then (λ i j, [chars.nth_le i (i.prop.trans_le h.1), chars.nth_le j (j.prop.trans_le h.2)].as_string) +else (λ i j, "_" ++ to_string i ++ "_" ++ to_string j) + +/-- `pi_fin.to_pexpr` but for matrices -/ +meta def fin_to_pexpr {m n : ℕ} (A : matrix (fin m) (fin n) pexpr) : pexpr := +``(@matrix.of (fin %%`(m)) (fin %%`(n)) _).app $ + pi_fin.to_pexpr (λ i : fin m, pi_fin.to_pexpr (λ j : fin n, A i j)) + +/-- This statement is defeq to `of_mul_of_fin`, but syntactically worse-/ +theorem of_mul_of_fin_aux (l m n : ℕ) ⦃α⦄ [has_mul α] [add_comm_monoid α] : + «forall» $ λ A : matrix (fin l) (fin m) α, + «forall» $ λ B : matrix (fin m) (fin n) α, + A.mul B = A.mulᵣ B := +by simp_rw [forall_iff, mulᵣ_eq, eq_self_iff_true, forall_const] + +/-- Prove a statement of the form +``` +∀ α [has_mul α] [add_comm_monoid α] (a₁₁ ... aₗₘ b₁₁ ... bₘₙ : α), + !![a₁₁ ⋱ aₗₘ] ⬝ !![b₁₁ ⋱ bₘₙ] = !![⋱] +``` +Returns the type of this statement and its proof. -/ +meta def of_mul_of_fin.prove (l m n : ℕ) : tactic (expr × expr) := +do + -- create all the binders, one for each coefficient + u ← tactic.mk_meta_univ, + α ← tactic.mk_local' `α binder_info.implicit (expr.sort u.succ), + has_mul_α ← tactic.mk_app `has_mul [α] >>= tactic.mk_local' `_inst_1 binder_info.inst_implicit, + add_comm_monoid_α ← + tactic.mk_app `add_comm_monoid [α] >>= tactic.mk_local' `_inst_2 binder_info.inst_implicit, + a ← (fin.mmap $ λ i : fin l, fin.mmap $ λ j : fin m, + tactic.mk_local' ((`a).append_suffix (name_suffix i j)) binder_info.default α), + b ← (fin.mmap $ λ i : fin m, fin.mmap $ λ j : fin n, + tactic.mk_local' ((`b).append_suffix (name_suffix i j)) binder_info.default α), + let a_flat := (list.fin_range l).bind (λ i, (list.fin_range m).map $ λ j, a i j), + let b_flat := (list.fin_range m).bind (λ i, (list.fin_range n).map $ λ j, b i j), + let args := [α, has_mul_α, add_comm_monoid_α] ++ a_flat ++ b_flat, + + -- build the matrices out of the coefficients + let A := matrix.fin_to_pexpr (matrix.map a to_pexpr), + let B := matrix.fin_to_pexpr (matrix.map b to_pexpr), + -- get an instance cache holding all the instances needed for matrix multiplication. There must + -- be a better way to do this. + t ← tactic.mk_instance_cache α, + has_add_α ← tactic.mk_app `has_add [α] >>= (λ t, prod.snd <$> @tactic.solve_aux unit t (do + { tmp2 ← tactic.pose `_inst_2' none add_comm_monoid_α, + tactic.reset_instance_cache, + tactic.apply_instance })), + has_zero_α ← tactic.mk_app `has_zero [α] >>= (λ t, prod.snd <$> @tactic.solve_aux unit t (do + { tmp2 ← tactic.pose `_inst_2' none add_comm_monoid_α, + tactic.reset_instance_cache, + tactic.apply_instance })), + let t := {inst := [ + (`has_mul, has_mul_α), + (`has_add, has_add_α), + (`has_zero, has_zero_α), + (`add_comm_monoid, add_comm_monoid_α)].foldl (λ n x, n.insert x.1 x.2) t.inst, + ..t}, + + -- clever trick: create algebraic instances on `expr` so that we can use `matrix.mul` or + -- `matrix.mulᵣ` to build the expression we want to end up with. It doesn't matter which we pick, + -- but the typeclasses are easier to create for the latter. + (t, has_mul_αe) ← expr.has_mul t, + (t, has_add_αe) ← expr.has_add t, + (t, has_zero_αe) ← expr.has_zero t, + let ab := @matrix.mulᵣ _ _ _ _ has_mul_αe has_add_αe has_zero_αe a b, + let AB := matrix.fin_to_pexpr (matrix.map ab to_pexpr), + + -- State and prove the equality, noting the RHS is defeq to `mulᵣ A B`. + A_eq ← tactic.to_expr ``(@matrix.mul _ _ _ _ _ %%has_mul_α %%add_comm_monoid_α %%A %%B = %%AB), + t ← tactic.pis args A_eq, + let pr := (expr.const `matrix.of_mul_of_fin_aux [u]).mk_app [`(l), `(m), `(n)], + -- This seems to create a metavariable then assign it, which ensures `pr` carries the right type. + ((), pr) ← tactic.solve_aux t $ tactic.exact pr, + + pure (t, pr) + +open_locale matrix + + +/-- Helper tactic used as an `auto_param` for `matrix.of_mul_of_fin` -/ +meta def of_mul_of_fin.derive : tactic unit := +do + target@`(@matrix.mul (fin %%l) (fin %%m) (fin %%n) %%α %%_ %%i1 %%i2 %%A %%B = %%AB) + ← tactic.target, + some (l, m, n) ← pure (prod.mk <$> l.to_nat <*> (prod.mk <$> m.to_nat <*> n.to_nat)) | + fail!"Dimensions {l}, {m} {n} are not numerals", + (t,pr) ← of_mul_of_fin.prove l m n, + tactic.apply (pr α i1 i2) {}, + tactic.done + -- TODO: should we be extracting the coefficients manually so we can do a full invocation as + -- something like: + -- tactic.unify target (t.instantiate_pis [α, A']), + -- tactic.exact (pr α A') + +/-- This lemma assumes that `a_coeffs` and `b_coeffs` refer to expressions of the form +`![![x₀₀, x₀₁], ![x₁₀, x₁₁]]`. It then uses an `auto_param` to populate `ab_coeffs` with an +expression of the same form, containing the appropriate expressions in terms of `+`, `*`, `aᵢⱼ`, +and `bⱼₖ`. -/ +theorem of_mul_of_fin {α} [has_mul α] [add_comm_monoid α] {l m n : ℕ} + {a_coeffs : fin l → fin m → α} + {b_coeffs : fin m → fin n → α} + {ab_coeffs : fin l → fin n → α} + (h : of a_coeffs ⬝ of b_coeffs = of ab_coeffs . of_mul_of_fin.derive) : + of a_coeffs ⬝ of b_coeffs = of ab_coeffs := h + +end of_mul_of_fin + +end matrix diff --git a/test/matrix_reflection.lean b/test/matrix_reflection.lean index 331d0a5573bfc..a734b8979715a 100644 --- a/test/matrix_reflection.lean +++ b/test/matrix_reflection.lean @@ -1,3 +1,4 @@ +import data.matrix.auto import data.matrix.reflection variables {α: Type*} @@ -17,3 +18,23 @@ example [add_comm_monoid α] [has_mul α] a₂₁*b₁₁ + a₂₂*b₂₁ + a₂₃*b₃₁, a₂₁*b₁₂ + a₂₂*b₂₂ + a₂₃*b₃₂, a₂₁*b₁₃ + a₂₂*b₂₃ + a₂₃*b₃₃; a₃₁*b₁₁ + a₃₂*b₂₁ + a₃₃*b₃₁, a₃₁*b₁₂ + a₃₂*b₂₂ + a₃₃*b₃₂, a₃₁*b₁₃ + a₃₂*b₂₃ + a₃₃*b₃₃] := (matrix.mulᵣ_eq _ _).symm + + +example {α} [add_comm_monoid α] [has_mul α] (a₁₁ a₁₂ a₂₁ a₂₂ b₁₁ b₁₂ b₂₁ b₂₂ : α) : + !![a₁₁, a₁₂; + a₂₁, a₂₂] ⬝ !![b₁₁, b₁₂; + b₂₁, b₂₂] = !![a₁₁ * b₁₁ + a₁₂ * b₂₁, a₁₁ * b₁₂ + a₁₂ * b₂₂; + a₂₁ * b₁₁ + a₂₂ * b₂₁, a₂₁ * b₁₂ + a₂₂ * b₂₂] := +begin + rw of_mul_of_fin, +end + +example {α} [add_comm_monoid α] [has_mul α] (a₁₁ a₁₂ b₁₁ b₁₂ b₂₁ b₂₂ : α) : + !![a₁₁, a₁₂] ⬝ !![b₁₁, b₁₂; + b₂₁, b₂₂] = !![a₁₁ * b₁₁ + a₁₂ * b₂₁, a₁₁ * b₁₂ + a₁₂ * b₂₂;] := +begin + -- if we really need it, we can get the proof directly like this + of_mul_of_fin.prove 1 2 2 >>= function.uncurry (tactic.assertv `h), + specialize @h α _ _, + rw of_mul_of_fin +end From c1acdccd694b692db2619fff903e0e40de428169 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sat, 15 Jul 2023 01:08:20 +0000 Subject: [PATCH 1254/1291] fix(tactic/norm_num): Do not allow typos in simp lemmas if they are unused (#17981) Previously ```lean import tactic.norm_num example : 1 + 1 = 2 := by norm_num [nonsense] ``` was legal. More realistically, if a lemma was added to the simp set then renamed, `norm_num [old_name]` would keep working even if the lemma no longer existed. This only seemed to happen once in mathlib, in the tests. It feels a bit silly to build the lemma set only to discard it, but we already rebuild the same lemma set on every iteration of the `repeat1` --- src/tactic/norm_num.lean | 9 ++++++--- test/norm_num.lean | 2 ++ test/norm_num_ext.lean | 2 +- 3 files changed, 9 insertions(+), 4 deletions(-) diff --git a/src/tactic/norm_num.lean b/src/tactic/norm_num.lean index eaeb50dd01860..27753225ea3be 100644 --- a/src/tactic/norm_num.lean +++ b/src/tactic/norm_num.lean @@ -1360,9 +1360,12 @@ use `get_step` to get the default `norm_num` set and `derive.step` for the basic simplifications. -/ meta def tactic.norm_num (step : expr → tactic (expr × expr)) (hs : list simp_arg_type) (l : interactive.loc) : tactic unit := -repeat1 $ orelse' (tactic.norm_num1 step l) $ -interactive.simp_core {} (tactic.norm_num1 step (interactive.loc.ns [none])) - ff (simp_arg_type.except ``one_div :: hs) [] l >> skip +do + -- Build and discard the simp lemma set, to validate it. + mk_simp_set_core ff [] (simp_arg_type.except ``one_div :: hs) tt, + repeat1 $ orelse' (tactic.norm_num1 step l) $ + interactive.simp_core {} (tactic.norm_num1 step (interactive.loc.ns [none])) + ff (simp_arg_type.except ``one_div :: hs) [] l >> skip /-- Carry out similar operations as `tactic.norm_num` but on an `expr` rather than a location. Given an expression `e`, returns `(e', ⊢ e = e')`. diff --git a/test/norm_num.lean b/test/norm_num.lean index d2b98fdec5b84..97dba721cd51b 100644 --- a/test/norm_num.lean +++ b/test/norm_num.lean @@ -304,6 +304,8 @@ example : (- ((- (((66 - 86) - 36) / 94) - 3) / - - (77 / (56 - - - 79))) + 87) example : 2 ^ 13 - 1 = int.of_nat 8191 := by norm_num +example : 1 + 1 = 2 := by success_if_fail { norm_num [this_doesnt_exist] }; refl + -- `^` and `•` do not have to match `monoid.has_pow` and `add_monoid.has_smul` syntactically example {α} [ring α] : (2 ^ 3 : ℕ → α) = 8 := by norm_num example {α} [ring α] : (2 • 3 : ℕ → α) = 6 := by norm_num diff --git a/test/norm_num_ext.lean b/test/norm_num_ext.lean index 96648a919b00d..5f75b9d9e5751 100644 --- a/test/norm_num_ext.lean +++ b/test/norm_num_ext.lean @@ -296,7 +296,7 @@ open_locale big_operators example : ([1, 2, 1, 3]).sum = 7 := by norm_num [-list.sum_cons] example : (([1, 2, 1, 3] : list ℚ).map (λ i, i^2)).sum = 15 := by norm_num [-list.map] example : (list.range 10).sum = 45 := by norm_num [-list.range_succ] -example : (list.fin_range 10).sum = 45 := by norm_num [-list.fin_range_succ] +example : (list.fin_range 10).sum = 45 := by norm_num [-list.fin_range_succ_eq_map] -- Multisets: example : (1 ::ₘ 2 ::ₘ 1 ::ₘ 3 ::ₘ {}).sum = 7 := by norm_num [-multiset.sum_cons] From 2fe465deb81bcd7ccafa065bb686888a82f15372 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Sat, 15 Jul 2023 12:19:30 +0000 Subject: [PATCH 1255/1291] chore(*): add mathlib4 synchronization comments (#19231) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.big_operators.norm_num` * `algebraic_geometry.projective_spectrum.scheme` * `analysis.analytic.inverse` * `analysis.inner_product_space.of_norm` * `analysis.normed.group.SemiNormedGroup.completion` * `category_theory.monad.monadicity` * `data.buffer.parser.basic` * `data.buffer.parser.numeral` * `data.fintype.array` * `data.hash_map` * `geometry.manifold.algebra.left_invariant_derivation` * `geometry.manifold.cont_mdiff_mfderiv` * `geometry.manifold.vector_bundle.hom` * `linear_algebra.clifford_algebra.even_equiv` * `ring_theory.etale` * `ring_theory.kaehler` * `tactic.group` * `testing.slim_check.functions` --- src/algebra/big_operators/norm_num.lean | 3 +++ src/algebraic_geometry/projective_spectrum/scheme.lean | 3 +++ src/analysis/analytic/inverse.lean | 3 +++ src/analysis/inner_product_space/of_norm.lean | 3 +++ src/analysis/normed/group/SemiNormedGroup/completion.lean | 3 +++ src/category_theory/monad/monadicity.lean | 3 +++ src/data/buffer/parser/basic.lean | 3 +++ src/data/buffer/parser/numeral.lean | 3 +++ src/data/fintype/array.lean | 3 +++ src/data/hash_map.lean | 3 +++ src/geometry/manifold/algebra/left_invariant_derivation.lean | 3 +++ src/geometry/manifold/cont_mdiff_mfderiv.lean | 3 +++ src/geometry/manifold/vector_bundle/hom.lean | 3 +++ src/linear_algebra/clifford_algebra/even_equiv.lean | 3 +++ src/ring_theory/etale.lean | 3 +++ src/ring_theory/kaehler.lean | 3 +++ src/tactic/group.lean | 3 +++ src/testing/slim_check/functions.lean | 3 +++ 18 files changed, 54 insertions(+) diff --git a/src/algebra/big_operators/norm_num.lean b/src/algebra/big_operators/norm_num.lean index a4067b58b3425..d950f11219ca6 100644 --- a/src/algebra/big_operators/norm_num.lean +++ b/src/algebra/big_operators/norm_num.lean @@ -9,6 +9,9 @@ import tactic.norm_num /-! ### `norm_num` plugin for big operators +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This `norm_num` plugin provides support for computing sums and products of lists, multisets and finsets. diff --git a/src/algebraic_geometry/projective_spectrum/scheme.lean b/src/algebraic_geometry/projective_spectrum/scheme.lean index 4bfbcf04dcd7e..4fc52b338df17 100644 --- a/src/algebraic_geometry/projective_spectrum/scheme.lean +++ b/src/algebraic_geometry/projective_spectrum/scheme.lean @@ -10,6 +10,9 @@ import ring_theory.graded_algebra.radical /-! # Proj as a scheme +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file is to prove that `Proj` is a scheme. ## Notation diff --git a/src/analysis/analytic/inverse.lean b/src/analysis/analytic/inverse.lean index 60268c17f70fd..d2c1a11a873ea 100644 --- a/src/analysis/analytic/inverse.lean +++ b/src/analysis/analytic/inverse.lean @@ -9,6 +9,9 @@ import tactic.congrm # Inverse of analytic functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We construct the left and right inverse of a formal multilinear series with invertible linear term, we prove that they coincide and study their properties (notably convergence). diff --git a/src/analysis/inner_product_space/of_norm.lean b/src/analysis/inner_product_space/of_norm.lean index be0bf32167923..7c36dba00df14 100644 --- a/src/analysis/inner_product_space/of_norm.lean +++ b/src/analysis/inner_product_space/of_norm.lean @@ -10,6 +10,9 @@ import analysis.inner_product_space.basic /-! # Inner product space derived from a norm +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines an `inner_product_space` instance from a norm that respects the parallellogram identity. The parallelogram identity is a way to express the inner product of `x` and `y` in terms of the norms of `x`, `y`, `x + y`, `x - y`. diff --git a/src/analysis/normed/group/SemiNormedGroup/completion.lean b/src/analysis/normed/group/SemiNormedGroup/completion.lean index 27a0a2315624c..8ffe9b3cff21e 100644 --- a/src/analysis/normed/group/SemiNormedGroup/completion.lean +++ b/src/analysis/normed/group/SemiNormedGroup/completion.lean @@ -10,6 +10,9 @@ import analysis.normed.group.hom_completion /-! # Completions of normed groups +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains an API for completions of seminormed groups (basic facts about objects and morphisms). diff --git a/src/category_theory/monad/monadicity.lean b/src/category_theory/monad/monadicity.lean index 8a15ac2fb4247..3caeaa2ad27b2 100644 --- a/src/category_theory/monad/monadicity.lean +++ b/src/category_theory/monad/monadicity.lean @@ -12,6 +12,9 @@ import category_theory.monad.limits /-! # Monadicity theorems +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We prove monadicity theorems which can establish a given functor is monadic. In particular, we show three versions of Beck's monadicity theorem, and the reflexive (crude) monadicity theorem: diff --git a/src/data/buffer/parser/basic.lean b/src/data/buffer/parser/basic.lean index 4c9b11621f64b..30f1b1f17d5ef 100644 --- a/src/data/buffer/parser/basic.lean +++ b/src/data/buffer/parser/basic.lean @@ -11,6 +11,9 @@ import data.buffer.parser /-! # Parsers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + `parser α` is the type that describes a computation that can ingest a `char_buffer` and output, if successful, a term of type `α`. This file expands on the definitions in the core library, proving that all the core library diff --git a/src/data/buffer/parser/numeral.lean b/src/data/buffer/parser/numeral.lean index af91eeba878d2..050b505ee8b8f 100644 --- a/src/data/buffer/parser/numeral.lean +++ b/src/data/buffer/parser/numeral.lean @@ -8,6 +8,9 @@ import data.buffer.parser.basic /-! # Numeral parsers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file expands on the existing `nat : parser ℕ` to provide parsers into any type `α` that can be represented by a numeral, which relies on `α` having a 0, 1, and addition operation. There are also convenience parsers that ensure that the numeral parsed in is not larger than diff --git a/src/data/fintype/array.lean b/src/data/fintype/array.lean index 5211c8cebc957..134d9d2837d71 100644 --- a/src/data/fintype/array.lean +++ b/src/data/fintype/array.lean @@ -8,6 +8,9 @@ import logic.equiv.array /-! # `array n α` is a fintype when `α` is. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ variables {α : Type*} diff --git a/src/data/hash_map.lean b/src/data/hash_map.lean index e9293d372477a..1c105b600c392 100644 --- a/src/data/hash_map.lean +++ b/src/data/hash_map.lean @@ -12,6 +12,9 @@ import data.pnat.defs /-! # Hash maps +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Defines a hash map data structure, representing a finite key-value map with a value type that may depend on the key type. The structure requires a `nat`-valued hash function to associate keys to buckets. diff --git a/src/geometry/manifold/algebra/left_invariant_derivation.lean b/src/geometry/manifold/algebra/left_invariant_derivation.lean index 6aab779498d14..eb5dd7aff2692 100644 --- a/src/geometry/manifold/algebra/left_invariant_derivation.lean +++ b/src/geometry/manifold/algebra/left_invariant_derivation.lean @@ -11,6 +11,9 @@ import geometry.manifold.derivation_bundle # Left invariant derivations +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define the concept of left invariant derivation for a Lie group. The concept is analogous to the more classical concept of left invariant vector fields, and it holds that the derivation associated to a vector field is left invariant iff the field is. diff --git a/src/geometry/manifold/cont_mdiff_mfderiv.lean b/src/geometry/manifold/cont_mdiff_mfderiv.lean index 9efd7835a3bd5..89abe55046496 100644 --- a/src/geometry/manifold/cont_mdiff_mfderiv.lean +++ b/src/geometry/manifold/cont_mdiff_mfderiv.lean @@ -9,6 +9,9 @@ import geometry.manifold.cont_mdiff_map /-! ### Interactions between differentiability, smoothness and manifold derivatives +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We give the relation between `mdifferentiable`, `cont_mdiff`, `mfderiv`, `tangent_map` and related notions. diff --git a/src/geometry/manifold/vector_bundle/hom.lean b/src/geometry/manifold/vector_bundle/hom.lean index 9dff42177cf64..ef6b4079edcfb 100644 --- a/src/geometry/manifold/vector_bundle/hom.lean +++ b/src/geometry/manifold/vector_bundle/hom.lean @@ -8,6 +8,9 @@ import topology.vector_bundle.hom /-! # Homs of smooth vector bundles over the same base space +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Here we show that `bundle.continuous_linear_map` is a smooth vector bundle. Note that we only do this for bundles of linear maps, not for bundles of arbitrary semilinear maps. diff --git a/src/linear_algebra/clifford_algebra/even_equiv.lean b/src/linear_algebra/clifford_algebra/even_equiv.lean index caf1ccefbe519..55e8ea60c4c02 100644 --- a/src/linear_algebra/clifford_algebra/even_equiv.lean +++ b/src/linear_algebra/clifford_algebra/even_equiv.lean @@ -9,6 +9,9 @@ import linear_algebra.quadratic_form.prod /-! # Isomorphisms with the even subalgebra of a Clifford algebra +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file provides some notable isomorphisms regarding the even subalgebra, `clifford_algebra.even`. ## Main definitions diff --git a/src/ring_theory/etale.lean b/src/ring_theory/etale.lean index 09b4302def2a3..93367714fc187 100644 --- a/src/ring_theory/etale.lean +++ b/src/ring_theory/etale.lean @@ -10,6 +10,9 @@ import ring_theory.kaehler # Formally étale morphisms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + An `R`-algebra `A` is formally étale (resp. unramified, smooth) if for every `R`-algebra, every square-zero ideal `I : ideal B` and `f : A →ₐ[R] B ⧸ I`, there exists exactly (resp. at most, at least) one lift `A →ₐ[R] B`. diff --git a/src/ring_theory/kaehler.lean b/src/ring_theory/kaehler.lean index 14eedac487b93..ec3b0a3079dc8 100644 --- a/src/ring_theory/kaehler.lean +++ b/src/ring_theory/kaehler.lean @@ -11,6 +11,9 @@ import ring_theory.is_tensor_product /-! # The module of kaehler differentials +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main results - `kaehler_differential`: The module of kaehler differentials. For an `R`-algebra `S`, we provide diff --git a/src/tactic/group.lean b/src/tactic/group.lean index 1cfe207169a8e..e9603a022658f 100644 --- a/src/tactic/group.lean +++ b/src/tactic/group.lean @@ -10,6 +10,9 @@ import algebra.group.commutator /-! # `group` +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Normalizes expressions in the language of groups. The basic idea is to use the simplifier to put everything into a product of group powers (`zpow` which takes a group element and an integer), then simplify the exponents using the `ring` tactic. The process needs to be repeated diff --git a/src/testing/slim_check/functions.lean b/src/testing/slim_check/functions.lean index b70e3757746c6..7249c61434dfe 100644 --- a/src/testing/slim_check/functions.lean +++ b/src/testing/slim_check/functions.lean @@ -14,6 +14,9 @@ import testing.slim_check.testable /-! ## `slim_check`: generators for functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines `sampleable` instances for `α → β` functions and `ℤ → ℤ` injective functions. From 31b269b60935483943542d547a6dd83a66b37dc7 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Sun, 16 Jul 2023 07:50:31 +0000 Subject: [PATCH 1256/1291] feat(set_theory/ordinal/natural_ops): define natural multiplication (#14324) --- src/set_theory/ordinal/arithmetic.lean | 14 + src/set_theory/ordinal/natural_ops.lean | 352 ++++++++++++++++++++++-- src/set_theory/ordinal/principal.lean | 22 +- 3 files changed, 355 insertions(+), 33 deletions(-) diff --git a/src/set_theory/ordinal/arithmetic.lean b/src/set_theory/ordinal/arithmetic.lean index a860ab829134a..daeab4abbea76 100644 --- a/src/set_theory/ordinal/arithmetic.lean +++ b/src/set_theory/ordinal/arithmetic.lean @@ -1526,6 +1526,20 @@ theorem is_normal.eq_iff_zero_and_succ {f g : ordinal.{u} → ordinal.{u}} (hf : exact H b hb end)⟩ +/-- A two-argument version of `ordinal.blsub`. + +We don't develop a full API for this, since it's only used in a handful of existence results. -/ +def blsub₂ (o₁ o₂ : ordinal) (op : Π (a < o₁) (b < o₂), ordinal) : ordinal := +lsub (λ x : o₁.out.α × o₂.out.α, + op (typein (<) x.1) (typein_lt_self _) (typein (<) x.2) (typein_lt_self _)) + +theorem lt_blsub₂ {o₁ o₂ : ordinal} (op : Π (a < o₁) (b < o₂), ordinal) {a b : ordinal} + (ha : a < o₁) (hb : b < o₂) : op a ha b hb < blsub₂ o₁ o₂ op := +begin + convert lt_lsub _ (prod.mk (enum (<) a (by rwa type_lt)) (enum (<) b (by rwa type_lt))), + simp only [typein_enum] +end + /-! ### Minimum excluded ordinals -/ /-- The minimum excluded ordinal in a family of ordinals. -/ diff --git a/src/set_theory/ordinal/natural_ops.lean b/src/set_theory/ordinal/natural_ops.lean index ff6f1e849beff..38f6a566ec06e 100644 --- a/src/set_theory/ordinal/natural_ops.lean +++ b/src/set_theory/ordinal/natural_ops.lean @@ -5,6 +5,7 @@ Authors: Violeta Hernández Palacios -/ import set_theory.ordinal.arithmetic +import tactic.abel /-! # Natural operations on ordinals @@ -38,7 +39,6 @@ between both types, we attempt to prove and state most results on `ordinal`. # Todo -- Define natural multiplication and provide a basic API. - Prove the characterizations of natural addition and multiplication in terms of the Cantor normal form. -/ @@ -49,6 +49,8 @@ open function order noncomputable theory +/-! ### Basic casts between `ordinal` and `nat_ordinal` -/ + /-- A type synonym for ordinals with natural addition and multiplication. -/ @[derive [has_zero, inhabited, has_one, linear_order, succ_order, has_well_founded]] def nat_ordinal : Type* := ordinal @@ -59,10 +61,10 @@ def nat_ordinal : Type* := ordinal /-- The identity function between `nat_ordinal` and `ordinal`. -/ @[pattern] def nat_ordinal.to_ordinal : nat_ordinal ≃o ordinal := order_iso.refl _ -open ordinal - namespace nat_ordinal +open ordinal + variables {a b c : nat_ordinal.{u}} @[simp] theorem to_ordinal_symm_eq : nat_ordinal.to_ordinal.symm = ordinal.to_nat_ordinal := rfl @@ -95,8 +97,6 @@ end nat_ordinal namespace ordinal -variables {a b c : ordinal.{u}} - @[simp] theorem to_nat_ordinal_symm_eq : to_nat_ordinal.symm = nat_ordinal.to_ordinal := rfl @[simp] theorem to_nat_ordinal_to_ordinal (a : ordinal) : a.to_nat_ordinal.to_ordinal = a := rfl @@ -106,11 +106,14 @@ variables {a b c : ordinal.{u}} @[simp] theorem to_nat_ordinal_eq_zero (a) : to_nat_ordinal a = 0 ↔ a = 0 := iff.rfl @[simp] theorem to_nat_ordinal_eq_one (a) : to_nat_ordinal a = 1 ↔ a = 1 := iff.rfl -@[simp] theorem to_nat_ordinal_max : +@[simp] theorem to_nat_ordinal_max (a b : ordinal) : to_nat_ordinal (max a b) = max a.to_nat_ordinal b.to_nat_ordinal := rfl -@[simp] theorem to_nat_ordinal_min : +@[simp] theorem to_nat_ordinal_min (a b : ordinal) : (linear_order.min a b).to_nat_ordinal = linear_order.min a.to_nat_ordinal b.to_nat_ordinal := rfl +/-! We place the definitions of `nadd` and `nmul` before actually developing their API, as this +guarantees we only need to open the `natural_ops` locale once. -/ + /-- Natural addition on ordinals `a ♯ b`, also known as the Hessenberg sum, is recursively defined as the least ordinal greater than `a' ♯ b` and `a ♯ b'` for all `a' < a` and `b' < b`. In contrast to normal ordinal addition, it is commutative. @@ -125,6 +128,30 @@ using_well_founded { dec_tac := `[solve_by_elim [psigma.lex.left, psigma.lex.rig localized "infix (name := ordinal.nadd) ` ♯ `:65 := ordinal.nadd" in natural_ops +/-- Natural multiplication on ordinals `a ⨳ b`, also known as the Hessenberg product, is recursively +defined as the least ordinal such that `a ⨳ b + a' ⨳ b'` is greater than `a' ⨳ b + a ⨳ b'` for all +`a' < a` and `b < b'`. In contrast to normal ordinal multiplication, it is commutative and +distributive (over natural addition). + +Natural multiplication can equivalently be characterized as the ordinal resulting from multiplying +the Cantor normal forms of `a` and `b` as if they were polynomials in `ω`. Addition of exponents is +done via natural addition. -/ +noncomputable def nmul : ordinal.{u} → ordinal.{u} → ordinal.{u} +| a b := Inf {c | ∀ (a' < a) (b' < b), nmul a' b ♯ nmul a b' < c ♯ nmul a' b'} +using_well_founded { dec_tac := `[solve_by_elim [psigma.lex.left, psigma.lex.right]] } + +localized "infix ` ⨳ `:70 := ordinal.nmul" in natural_ops + +end ordinal + +open_locale natural_ops + +/-! ### Natural addition -/ + +namespace ordinal + +variables {a b c : ordinal.{u}} + theorem nadd_def (a b : ordinal) : a ♯ b = max (blsub.{u u} a $ λ a' h, a' ♯ b) (blsub.{u u} b $ λ b' h, a ♯ b') := @@ -243,10 +270,10 @@ end end ordinal -open ordinal - namespace nat_ordinal +open ordinal + instance : has_add nat_ordinal := ⟨nadd⟩ instance add_covariant_class_lt : @@ -280,21 +307,19 @@ instance : add_monoid_with_one nat_ordinal := add_monoid_with_one.unary begin induction n with n hn, { refl }, - { change nadd (to_ordinal n) 1 = n + 1, - rw hn, - apply nadd_one } + { change to_ordinal n ♯ 1 = n + 1, + rw hn, exact nadd_one n } end end nat_ordinal -open nat_ordinal - -open_locale natural_ops - namespace ordinal +theorem nadd_eq_add (a b : ordinal) : a ♯ b = (a.to_nat_ordinal + b.to_nat_ordinal).to_ordinal := +rfl + @[simp] theorem to_nat_ordinal_cast_nat (n : ℕ) : to_nat_ordinal n = n := -by { rw ←to_ordinal_cast_nat n, refl } +by { rw ←nat_ordinal.to_ordinal_cast_nat n, refl } theorem lt_of_nadd_lt_nadd_left : ∀ {a b c}, a ♯ b < a ♯ c → b < c := @lt_of_add_lt_add_left nat_ordinal _ _ _ @@ -314,6 +339,16 @@ theorem nadd_le_nadd_iff_left : ∀ a {b c}, a ♯ b ≤ a ♯ c ↔ b ≤ c := theorem nadd_le_nadd_iff_right : ∀ a {b c}, b ♯ a ≤ c ♯ a ↔ b ≤ c := @_root_.add_le_add_iff_right nat_ordinal _ _ _ _ +theorem nadd_le_nadd : ∀ {a b c d}, a ≤ b → c ≤ d → a ♯ c ≤ b ♯ d := +@add_le_add nat_ordinal _ _ _ _ +theorem nadd_lt_nadd : ∀ {a b c d}, a < b → c < d → a ♯ c < b ♯ d := +@add_lt_add nat_ordinal _ _ _ _ + +theorem nadd_lt_nadd_of_lt_of_le : ∀ {a b c d}, a < b → c ≤ d → a ♯ c < b ♯ d := +@add_lt_add_of_lt_of_le nat_ordinal _ _ _ _ +theorem nadd_lt_nadd_of_le_of_lt : ∀ {a b c d}, a ≤ b → c < d → a ♯ c < b ♯ d := +@add_lt_add_of_le_of_lt nat_ordinal _ _ _ _ + theorem nadd_left_cancel : ∀ {a b c}, a ♯ b = a ♯ c → b = c := @_root_.add_left_cancel nat_ordinal _ _ theorem nadd_right_cancel : ∀ {a b c}, a ♯ b = c ♯ b → a = c := @@ -323,4 +358,287 @@ theorem nadd_left_cancel_iff : ∀ {a b c}, a ♯ b = a ♯ c ↔ b = c := theorem nadd_right_cancel_iff : ∀ {a b c}, b ♯ a = c ♯ a ↔ b = c := @add_right_cancel_iff nat_ordinal _ _ +theorem le_nadd_self {a b} : a ≤ b ♯ a := +by simpa using nadd_le_nadd_right (ordinal.zero_le b) a +theorem le_nadd_left {a b c} (h : a ≤ c) : a ≤ b ♯ c := +le_nadd_self.trans (nadd_le_nadd_left h b) +theorem le_self_nadd {a b} : a ≤ a ♯ b := +by simpa using nadd_le_nadd_left (ordinal.zero_le b) a +theorem le_nadd_right {a b c} (h : a ≤ b) : a ≤ b ♯ c := +le_self_nadd.trans (nadd_le_nadd_right h c) + +theorem nadd_left_comm : ∀ a b c, a ♯ (b ♯ c) = b ♯ (a ♯ c) := +@add_left_comm nat_ordinal _ +theorem nadd_right_comm : ∀ a b c, a ♯ b ♯ c = a ♯ c ♯ b := +@add_right_comm nat_ordinal _ + +/-! ### Natural multiplication -/ + +variables {a b c d : ordinal.{u}} + +theorem nmul_def (a b : ordinal) : + a ⨳ b = Inf {c | ∀ (a' < a) (b' < b), a' ⨳ b ♯ a ⨳ b' < c ♯ a' ⨳ b'} := +by rw nmul + +/-- The set in the definition of `nmul` is nonempty. -/ +theorem nmul_nonempty (a b : ordinal.{u}) : + {c : ordinal.{u} | ∀ (a' < a) (b' < b), a' ⨳ b ♯ a ⨳ b' < c ♯ a' ⨳ b'}.nonempty := +⟨_, λ a' ha b' hb, (lt_blsub₂.{u u u} _ ha hb).trans_le le_self_nadd⟩ + +theorem nmul_nadd_lt {a' b' : ordinal} (ha : a' < a) (hb : b' < b) : + a' ⨳ b ♯ a ⨳ b' < a ⨳ b ♯ a' ⨳ b' := +by { rw nmul_def a b, exact Inf_mem (nmul_nonempty a b) a' ha b' hb } + +theorem nmul_nadd_le {a' b' : ordinal} (ha : a' ≤ a) (hb : b' ≤ b) : + a' ⨳ b ♯ a ⨳ b' ≤ a ⨳ b ♯ a' ⨳ b' := +begin + rcases lt_or_eq_of_le ha with ha | rfl, + { rcases lt_or_eq_of_le hb with hb | rfl, + { exact (nmul_nadd_lt ha hb).le }, + { rw nadd_comm } }, + { exact le_rfl } +end + +theorem lt_nmul_iff : c < a ⨳ b ↔ ∃ (a' < a) (b' < b), c ♯ a' ⨳ b' ≤ a' ⨳ b ♯ a ⨳ b' := +begin + refine ⟨λ h, _, _⟩, + { rw nmul at h, + simpa using not_mem_of_lt_cInf h ⟨0, λ _ _, bot_le⟩ }, + { rintros ⟨a', ha, b', hb, h⟩, + have := h.trans_lt (nmul_nadd_lt ha hb), + rwa nadd_lt_nadd_iff_right at this } +end + +theorem nmul_le_iff : a ⨳ b ≤ c ↔ ∀ (a' < a) (b' < b), a' ⨳ b ♯ a ⨳ b' < c ♯ a' ⨳ b' := +by { rw ←not_iff_not, simp [lt_nmul_iff] } + +theorem nmul_comm : ∀ (a b), a ⨳ b = b ⨳ a +| a b := begin + rw [nmul, nmul], + congr, ext x, split; + intros H c hc d hd, + { rw [nadd_comm, ←nmul_comm, ←nmul_comm a, ←nmul_comm d], + exact H _ hd _ hc }, + { rw [nadd_comm, nmul_comm, nmul_comm c, nmul_comm c], + exact H _ hd _ hc } +end +using_well_founded { dec_tac := `[solve_by_elim [psigma.lex.left, psigma.lex.right]] } + +@[simp] theorem nmul_zero (a) : a ⨳ 0 = 0 := +by { rw [←ordinal.le_zero, nmul_le_iff], exact λ _ _ a ha, (ordinal.not_lt_zero a ha).elim } + +@[simp] theorem zero_nmul (a) : 0 ⨳ a = 0 := +by rw [nmul_comm, nmul_zero] + +@[simp] theorem nmul_one : ∀ a, a ⨳ 1 = a +| a := begin + rw nmul, + simp only [lt_one_iff_zero, forall_eq, nmul_zero, nadd_zero], + convert cInf_Ici, + ext b, + refine ⟨λ H, le_of_forall_lt (λ c hc, _), λ ha c hc, _⟩, + { simpa only [nmul_one] using H c hc }, + { simpa only [nmul_one] using hc.trans_le ha } +end +using_well_founded { dec_tac := `[assumption] } + +@[simp] theorem one_nmul (a) : 1 ⨳ a = a := +by rw [nmul_comm, nmul_one] + +theorem nmul_lt_nmul_of_pos_left (h₁ : a < b) (h₂ : 0 < c) : c ⨳ a < c ⨳ b := +lt_nmul_iff.2 ⟨0, h₂, a, h₁, by simp⟩ + +theorem nmul_lt_nmul_of_pos_right (h₁ : a < b) (h₂ : 0 < c) : a ⨳ c < b ⨳ c := +lt_nmul_iff.2 ⟨a, h₁, 0, h₂, by simp⟩ + +theorem nmul_le_nmul_of_nonneg_left (h₁ : a ≤ b) (h₂ : 0 ≤ c) : c ⨳ a ≤ c ⨳ b := +begin + rcases lt_or_eq_of_le h₁ with h₁|rfl; + rcases lt_or_eq_of_le h₂ with h₂|rfl, + { exact (nmul_lt_nmul_of_pos_left h₁ h₂).le }, + all_goals { simp } +end + +theorem nmul_le_nmul_of_nonneg_right (h₁ : a ≤ b) (h₂ : 0 ≤ c) : a ⨳ c ≤ b ⨳ c := +begin + rw [nmul_comm, nmul_comm b], + exact nmul_le_nmul_of_nonneg_left h₁ h₂ +end + +theorem nmul_nadd : ∀ (a b c), a ⨳ (b ♯ c) = a ⨳ b ♯ a ⨳ c +| a b c := begin + apply le_antisymm (nmul_le_iff.2 $ λ a' ha d hd, _) (nadd_le_iff.2 ⟨λ d hd, _, λ d hd, _⟩), + { rw nmul_nadd, + rcases lt_nadd_iff.1 hd with ⟨b', hb, hd⟩ | ⟨c', hc, hd⟩, + { have := nadd_lt_nadd_of_lt_of_le (nmul_nadd_lt ha hb) (nmul_nadd_le ha.le hd), + rw [nmul_nadd, nmul_nadd] at this, + simp only [nadd_assoc] at this, + rwa [nadd_left_comm, nadd_left_comm _ (a ⨳ b'), nadd_left_comm (a ⨳ b), nadd_lt_nadd_iff_left, + nadd_left_comm (a' ⨳ b), nadd_left_comm (a ⨳ b), nadd_lt_nadd_iff_left, ←nadd_assoc, + ←nadd_assoc] at this }, + { have := nadd_lt_nadd_of_le_of_lt (nmul_nadd_le ha.le hd) (nmul_nadd_lt ha hc), + rw [nmul_nadd, nmul_nadd] at this, + simp only [nadd_assoc] at this, + rwa [nadd_left_comm, nadd_comm (a ⨳ c), nadd_left_comm (a' ⨳ d), nadd_left_comm (a ⨳ c'), + nadd_left_comm (a ⨳ b), nadd_lt_nadd_iff_left, nadd_comm (a' ⨳ c), nadd_left_comm (a ⨳ d), + nadd_left_comm (a' ⨳ b), nadd_left_comm (a ⨳ b), nadd_lt_nadd_iff_left, nadd_comm (a ⨳ d), + nadd_comm (a' ⨳ d), ←nadd_assoc, ←nadd_assoc] at this } }, + { rcases lt_nmul_iff.1 hd with ⟨a', ha, b', hb, hd⟩, + have := nadd_lt_nadd_of_le_of_lt hd (nmul_nadd_lt ha (nadd_lt_nadd_right hb c)), + rw [nmul_nadd, nmul_nadd, nmul_nadd a'] at this, + simp only [nadd_assoc] at this, + rwa [nadd_left_comm (a' ⨳ b'), nadd_left_comm, nadd_lt_nadd_iff_left, nadd_left_comm, + nadd_left_comm _ (a' ⨳ b'), nadd_left_comm (a ⨳ b'), nadd_lt_nadd_iff_left, + nadd_left_comm (a' ⨳ c), nadd_left_comm, nadd_lt_nadd_iff_left, nadd_left_comm, + nadd_comm _ (a' ⨳ c), nadd_lt_nadd_iff_left] at this }, + { rcases lt_nmul_iff.1 hd with ⟨a', ha, c', hc, hd⟩, + have := nadd_lt_nadd_of_lt_of_le (nmul_nadd_lt ha (nadd_lt_nadd_left hc b)) hd, + rw [nmul_nadd, nmul_nadd, nmul_nadd a'] at this, + simp only [nadd_assoc] at this, + rwa [nadd_left_comm _ (a' ⨳ b), nadd_lt_nadd_iff_left, nadd_left_comm (a' ⨳ c'), + nadd_left_comm _ (a' ⨳ c), nadd_lt_nadd_iff_left, nadd_left_comm, + nadd_comm (a' ⨳ c'), nadd_left_comm _ (a ⨳ c'), nadd_lt_nadd_iff_left, + nadd_comm _ (a' ⨳ c'), nadd_comm _ (a' ⨳ c'), nadd_left_comm, + nadd_lt_nadd_iff_left] at this } +end +using_well_founded { dec_tac := `[solve_by_elim [psigma.lex.left, psigma.lex.right]] } + +theorem nadd_nmul (a b c) : (a ♯ b) ⨳ c = a ⨳ c ♯ b ⨳ c := +by rw [nmul_comm, nmul_nadd, nmul_comm, nmul_comm c] + +theorem nmul_nadd_lt₃ {a' b' c' : ordinal} (ha : a' < a) (hb : b' < b) (hc : c' < c) : + a' ⨳ b ⨳ c ♯ a ⨳ b' ⨳ c ♯ a ⨳ b ⨳ c' ♯ a' ⨳ b' ⨳ c' < + a ⨳ b ⨳ c ♯ a' ⨳ b' ⨳ c ♯ a' ⨳ b ⨳ c' ♯ a ⨳ b' ⨳ c' := +by simpa only [nadd_nmul, ←nadd_assoc] using nmul_nadd_lt (nmul_nadd_lt ha hb) hc + +theorem nmul_nadd_le₃ {a' b' c' : ordinal} (ha : a' ≤ a) (hb : b' ≤ b) (hc : c' ≤ c) : + a' ⨳ b ⨳ c ♯ a ⨳ b' ⨳ c ♯ a ⨳ b ⨳ c' ♯ a' ⨳ b' ⨳ c' ≤ + a ⨳ b ⨳ c ♯ a' ⨳ b' ⨳ c ♯ a' ⨳ b ⨳ c' ♯ a ⨳ b' ⨳ c' := +by simpa only [nadd_nmul, ←nadd_assoc] using nmul_nadd_le (nmul_nadd_le ha hb) hc + +theorem nmul_nadd_lt₃' {a' b' c' : ordinal} (ha : a' < a) (hb : b' < b) (hc : c' < c) : + a' ⨳ (b ⨳ c) ♯ a ⨳ (b' ⨳ c) ♯ a ⨳ (b ⨳ c') ♯ a' ⨳ (b' ⨳ c') < + a ⨳ (b ⨳ c) ♯ a' ⨳ (b' ⨳ c) ♯ a' ⨳ (b ⨳ c') ♯ a ⨳ (b' ⨳ c') := +begin + simp only [nmul_comm _ (_ ⨳ _)], + convert nmul_nadd_lt₃ hb hc ha using 1; + { simp only [nadd_eq_add, nat_ordinal.to_ordinal_to_nat_ordinal], abel } +end + +theorem nmul_nadd_le₃' {a' b' c' : ordinal} (ha : a' ≤ a) (hb : b' ≤ b) (hc : c' ≤ c) : + a' ⨳ (b ⨳ c) ♯ a ⨳ (b' ⨳ c) ♯ a ⨳ (b ⨳ c') ♯ a' ⨳ (b' ⨳ c') ≤ + a ⨳ (b ⨳ c) ♯ a' ⨳ (b' ⨳ c) ♯ a' ⨳ (b ⨳ c') ♯ a ⨳ (b' ⨳ c') := +begin + simp only [nmul_comm _ (_ ⨳ _)], + convert nmul_nadd_le₃ hb hc ha using 1; + { simp only [nadd_eq_add, nat_ordinal.to_ordinal_to_nat_ordinal], abel } +end + +theorem lt_nmul_iff₃ : d < a ⨳ b ⨳ c ↔ ∃ (a' < a) (b' < b) (c' < c), + d ♯ a' ⨳ b' ⨳ c ♯ a' ⨳ b ⨳ c' ♯ a ⨳ b' ⨳ c' ≤ + a' ⨳ b ⨳ c ♯ a ⨳ b' ⨳ c ♯ a ⨳ b ⨳ c' ♯ a' ⨳ b' ⨳ c' := +begin + refine ⟨λ h, _, _⟩, + { rcases lt_nmul_iff.1 h with ⟨e, he, c', hc, H₁⟩, + rcases lt_nmul_iff.1 he with ⟨a', ha, b', hb, H₂⟩, + refine ⟨a', ha, b', hb, c', hc, _⟩, + have := nadd_le_nadd H₁ (nmul_nadd_le H₂ hc.le), + simp only [nadd_nmul, nadd_assoc] at this, + rw [nadd_left_comm, nadd_left_comm d, nadd_left_comm, nadd_le_nadd_iff_left, + nadd_left_comm (a ⨳ b' ⨳ c), nadd_left_comm (a' ⨳ b ⨳ c), nadd_left_comm (a ⨳ b ⨳ c'), + nadd_le_nadd_iff_left, nadd_left_comm (a ⨳ b ⨳ c'), nadd_left_comm (a ⨳ b ⨳ c')] at this, + simpa only [nadd_assoc] }, + { rintro ⟨a', ha, b', hb, c', hc, h⟩, + have := h.trans_lt (nmul_nadd_lt₃ ha hb hc), + repeat { rwa nadd_lt_nadd_iff_right at this } } +end + +theorem nmul_le_iff₃ : a ⨳ b ⨳ c ≤ d ↔ ∀ (a' < a) (b' < b) (c' < c), + a' ⨳ b ⨳ c ♯ a ⨳ b' ⨳ c ♯ a ⨳ b ⨳ c' ♯ a' ⨳ b' ⨳ c' < + d ♯ a' ⨳ b' ⨳ c ♯ a' ⨳ b ⨳ c' ♯ a ⨳ b' ⨳ c' := +by { rw ←not_iff_not, simp [lt_nmul_iff₃] } + +theorem lt_nmul_iff₃' : d < a ⨳ (b ⨳ c) ↔ ∃ (a' < a) (b' < b) (c' < c), + d ♯ a' ⨳ (b' ⨳ c) ♯ a' ⨳ (b ⨳ c') ♯ a ⨳ (b' ⨳ c') ≤ + a' ⨳ (b ⨳ c) ♯ a ⨳ (b' ⨳ c) ♯ a ⨳ (b ⨳ c') ♯ a' ⨳ (b' ⨳ c') := +begin + simp only [nmul_comm _ (_ ⨳ _), lt_nmul_iff₃, nadd_eq_add, nat_ordinal.to_ordinal_to_nat_ordinal], + split; rintro ⟨b', hb, c', hc, a', ha, h⟩, + { use [a', ha, b', hb, c', hc], convert h using 1; abel }, + { use [c', hc, a', ha, b', hb], convert h using 1; abel } +end + +theorem nmul_le_iff₃' : a ⨳ (b ⨳ c) ≤ d ↔ ∀ (a' < a) (b' < b) (c' < c), + a' ⨳ (b ⨳ c) ♯ a ⨳ (b' ⨳ c) ♯ a ⨳ (b ⨳ c') ♯ a' ⨳ (b' ⨳ c') < + d ♯ a' ⨳ (b' ⨳ c) ♯ a' ⨳ (b ⨳ c') ♯ a ⨳ (b' ⨳ c') := +by { rw ←not_iff_not, simp [lt_nmul_iff₃'] } + +theorem nmul_assoc : ∀ a b c, a ⨳ b ⨳ c = a ⨳ (b ⨳ c) +| a b c := begin + apply le_antisymm, + { rw nmul_le_iff₃, + intros a' ha b' hb c' hc, + repeat { rw nmul_assoc }, + exact nmul_nadd_lt₃' ha hb hc }, + { rw nmul_le_iff₃', + intros a' ha b' hb c' hc, + repeat { rw ←nmul_assoc }, + exact nmul_nadd_lt₃ ha hb hc }, +end +using_well_founded { dec_tac := `[solve_by_elim [psigma.lex.left, psigma.lex.right]] } + +end ordinal + +open ordinal + +instance : has_mul nat_ordinal := ⟨nmul⟩ + +instance : ordered_comm_semiring nat_ordinal := +{ mul := (*), + left_distrib := nmul_nadd, + right_distrib := nadd_nmul, + zero_mul := zero_nmul, + mul_zero := nmul_zero, + mul_assoc := nmul_assoc, + one := 1, + one_mul := one_nmul, + mul_one := nmul_one, + mul_comm := nmul_comm, + zero_le_one := @zero_le_one ordinal _ _ _ _, + mul_le_mul_of_nonneg_left := λ a b c, nmul_le_nmul_of_nonneg_left, + mul_le_mul_of_nonneg_right := λ a b c, nmul_le_nmul_of_nonneg_right, + ..nat_ordinal.ordered_cancel_add_comm_monoid, + ..nat_ordinal.linear_order } + +namespace ordinal + +theorem nmul_eq_mul (a b) : a ⨳ b = (a.to_nat_ordinal * b.to_nat_ordinal).to_ordinal := rfl + +theorem nmul_nadd_one : ∀ a b, a ⨳ (b ♯ 1) = a ⨳ b ♯ a := @mul_add_one nat_ordinal _ _ _ +theorem nadd_one_nmul : ∀ a b, (a ♯ 1) ⨳ b = a ⨳ b ♯ b := @add_one_mul nat_ordinal _ _ _ +theorem nmul_succ (a b) : a ⨳ succ b = a ⨳ b ♯ a := by rw [←nadd_one, nmul_nadd_one] +theorem succ_nmul (a b) : succ a ⨳ b = a ⨳ b ♯ b := by rw [←nadd_one, nadd_one_nmul] +theorem nmul_add_one : ∀ a b, a ⨳ (b + 1) = a ⨳ b ♯ a := nmul_succ +theorem add_one_nmul : ∀ a b, (a + 1) ⨳ b = a ⨳ b ♯ b := succ_nmul + end ordinal + +namespace nat_ordinal + +open ordinal + +theorem mul_le_nmul (a b : ordinal.{u}) : a * b ≤ a ⨳ b := +begin + apply b.limit_rec_on, + { simp }, + { intros c h, + rw [mul_succ, nmul_succ], + exact (add_le_nadd _ a).trans (nadd_le_nadd_right h a) }, + { intros c hc H, + rcases eq_zero_or_pos a with rfl | ha, + { simp }, + { rw [←is_normal.blsub_eq.{u u} (mul_is_normal ha) hc, blsub_le_iff], + exact λ i hi, (H i hi).trans_lt (nmul_lt_nmul_of_pos_left hi ha) } } +end + +end nat_ordinal diff --git a/src/set_theory/ordinal/principal.lean b/src/set_theory/ordinal/principal.lean index 61774bc56030a..ce63acccbfca8 100644 --- a/src/set_theory/ordinal/principal.lean +++ b/src/set_theory/ordinal/principal.lean @@ -85,31 +85,21 @@ nfp_le $ λ n, (ho.iterate_lt hao n).le /-! ### Principal ordinals are unbounded -/ -/-- The least strict upper bound of `op` applied to all pairs of ordinals less than `o`. This is -essentially a two-argument version of `ordinal.blsub`. -/ -def blsub₂ (op : ordinal → ordinal → ordinal) (o : ordinal) : ordinal := -lsub (λ x : o.out.α × o.out.α, op (typein (<) x.1) (typein (<) x.2)) - -theorem lt_blsub₂ (op : ordinal → ordinal → ordinal) {o : ordinal} {a b : ordinal} (ha : a < o) - (hb : b < o) : op a b < blsub₂ op o := -begin - convert lt_lsub _ (prod.mk (enum (<) a (by rwa type_lt)) (enum (<) b (by rwa type_lt))), - simp only [typein_enum] -end - theorem principal_nfp_blsub₂ (op : ordinal → ordinal → ordinal) (o : ordinal) : - principal op (nfp (blsub₂.{u u} op) o) := + principal op (nfp (λ o', blsub₂.{u u u} o' o' (λ a _ b _, op a b)) o) := λ a b ha hb, begin rw lt_nfp at *, cases ha with m hm, cases hb with n hn, - cases le_total ((blsub₂.{u u} op)^[m] o) ((blsub₂.{u u} op)^[n] o) with h h, + cases le_total + ((λ o', blsub₂.{u u u} o' o' (λ a _ b _, op a b))^[m] o) + ((λ o', blsub₂.{u u u} o' o' (λ a _ b _, op a b))^[n] o) with h h, { use n + 1, rw function.iterate_succ', - exact lt_blsub₂ op (hm.trans_le h) hn }, + exact lt_blsub₂ _ (hm.trans_le h) hn }, { use m + 1, rw function.iterate_succ', - exact lt_blsub₂ op hm (hn.trans_le h) }, + exact lt_blsub₂ _ hm (hn.trans_le h) } end theorem unbounded_principal (op : ordinal → ordinal → ordinal) : From 016138c2e83fa76d338d5df7d32d0acb6c587792 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sun, 16 Jul 2023 09:47:03 +0000 Subject: [PATCH 1257/1291] fix(tactic/linarith): instantiate metavariables in linarith (#19233) This probably isn't exhaustive, but catches the really trivial case. This doesn't appear to need forward-porting. --- src/tactic/linarith/datatypes.lean | 4 ++-- src/tactic/linarith/frontend.lean | 6 +++--- src/tactic/linarith/parsing.lean | 4 ++-- src/tactic/linarith/preprocessing.lean | 18 +++++++++--------- src/tactic/linarith/verification.lean | 6 +++--- test/linarith.lean | 9 +++++++++ 6 files changed, 28 insertions(+), 19 deletions(-) diff --git a/src/tactic/linarith/datatypes.lean b/src/tactic/linarith/datatypes.lean index e1edff788486d..52bc606a08fb8 100644 --- a/src/tactic/linarith/datatypes.lean +++ b/src/tactic/linarith/datatypes.lean @@ -369,13 +369,13 @@ Typically `R` and `R'` will be the same, except when `c = 0`, in which case `R'` If `c = 1`, `h'` is the same as `h` -- specifically, it does *not* change the type to `1*t R 0`. -/ meta def mk_single_comp_zero_pf (c : ℕ) (h : expr) : tactic (ineq × expr) := -do tp ← infer_type h, +do tp ← infer_type h >>= instantiate_mvars, some (iq, e) ← return $ parse_into_comp_and_expr tp, if c = 0 then do e' ← mk_app ``zero_mul [e], return (ineq.eq, e') else if c = 1 then return (iq, h) else - do tp ← (prod.snd <$> (infer_type h >>= get_rel_sides)) >>= infer_type, + do tp ← (prod.snd <$> (infer_type h >>= instantiate_mvars >>= get_rel_sides)) >>= infer_type, c ← tp.of_nat c, cpos ← to_expr ``(%%c > 0), (_, ex) ← solve_aux cpos `[norm_num, done], diff --git a/src/tactic/linarith/frontend.lean b/src/tactic/linarith/frontend.lean index 5ea2c37df8738..f09b0cedc5ca5 100644 --- a/src/tactic/linarith/frontend.lean +++ b/src/tactic/linarith/frontend.lean @@ -151,7 +151,7 @@ newly introduced local constant. Otherwise returns `none`. -/ meta def apply_contr_lemma : tactic (option (expr × expr)) := -do t ← target, +do t ← target >>= instantiate_mvars, match get_contr_lemma_name_and_type t with | some (nm, tp) := do refine ((expr.const nm []) pexpr.mk_placeholder), v ← intro1, return $ some (tp, v) @@ -207,7 +207,7 @@ to only those that are comparisons over the type `restr_type`. -/ meta def filter_hyps_to_type (restr_type : expr) (hyps : list expr) : tactic (list expr) := hyps.mfilter $ λ h, do - ht ← infer_type h, + ht ← infer_type h >>= instantiate_mvars, match get_contr_lemma_name_and_type ht with | some (_, htype) := succeeds $ unify htype restr_type | none := return ff @@ -238,7 +238,7 @@ expressions. meta def tactic.linarith (reduce_semi : bool) (only_on : bool) (hyps : list pexpr) (cfg : linarith_config := {}) : tactic unit := focus1 $ -do t ← target, +do t ← target >>= instantiate_mvars, -- if the target is an equality, we run `linarith` twice, to prove ≤ and ≥. if t.is_eq.is_some then linarith_trace "target is an equality: splitting" >> diff --git a/src/tactic/linarith/parsing.lean b/src/tactic/linarith/parsing.lean index c31cdc3b9fcee..f43291cd087ea 100644 --- a/src/tactic/linarith/parsing.lean +++ b/src/tactic/linarith/parsing.lean @@ -182,8 +182,8 @@ It also returns the largest variable index that appears in comparisons in `c`. -/ meta def linear_forms_and_max_var (red : transparency) (pfs : list expr) : tactic (list comp × ℕ) := -do pftps ← pfs.mmap infer_type, - (l, _, map) ← to_comp_fold red [] pftps mk_rb_map, +do pftps ← pfs.mmap (λ e, infer_type e >>= instantiate_mvars), + (l, _, map) ← to_comp_fold red [] pftps mk_rb_map, return (l, map.size - 1) diff --git a/src/tactic/linarith/preprocessing.lean b/src/tactic/linarith/preprocessing.lean index 59775da603e99..ae700421d8eeb 100644 --- a/src/tactic/linarith/preprocessing.lean +++ b/src/tactic/linarith/preprocessing.lean @@ -70,7 +70,7 @@ private meta def rearr_comp_aux : expr → expr → tactic expr and turns it into a proof of a comparison `_ R 0`, where `R ∈ {=, ≤, <}`. -/ meta def rearr_comp (e : expr) : tactic expr := -infer_type e >>= rearr_comp_aux e +infer_type e >>= instantiate_mvars >>= rearr_comp_aux e /-- If `e` is of the form `((n : ℕ) : ℤ)`, `is_nat_int_coe e` returns `n : ℕ`. -/ meta def is_nat_int_coe : expr → option expr @@ -96,7 +96,7 @@ If `pf` is a proof of a strict inequality `(a : ℤ) < b`, and similarly if `pf` proves a negated weak inequality. -/ meta def mk_non_strict_int_pf_of_strict_int_pf (pf : expr) : tactic expr := -do tp ← infer_type pf, +do tp ← infer_type pf >>= instantiate_mvars, match tp with | `(%%a < %%b) := to_expr ``(int.add_one_le_iff.mpr %%pf) | `(%%a > %%b) := to_expr ``(int.add_one_le_iff.mpr %%pf) @@ -139,7 +139,7 @@ Removes any expressions that are not proofs of inequalities, equalities, or nega meta def filter_comparisons : preprocessor := { name := "filter terms that are not proofs of comparisons", transform := λ h, -(do tp ← infer_type h, +(do tp ← infer_type h >>= instantiate_mvars, is_prop tp >>= guardb, guardb (filter_comparisons_aux tp), return [h]) @@ -152,7 +152,7 @@ For example, a proof of `¬ a < b` will become a proof of `a ≥ b`. meta def remove_negations : preprocessor := { name := "replace negations of comparisons", transform := λ h, -do tp ← infer_type h, +do tp ← infer_type h >>= instantiate_mvars, match tp with | `(¬ %%p) := singleton <$> rem_neg h p | _ := return [h] @@ -171,9 +171,9 @@ meta def nat_to_int : global_preprocessor := -- we lock the tactic state here because a `simplify` call inside of -- `zify_proof` corrupts the tactic state when run under `io.run_tactic`. do l ← lock_tactic_state $ l.mmap $ λ h, - infer_type h >>= guardb ∘ is_nat_prop >> zify_proof [] h <|> return h, + infer_type h >>= instantiate_mvars >>= guardb ∘ is_nat_prop >> zify_proof [] h <|> return h, nonnegs ← l.mfoldl (λ (es : expr_set) h, do - (a, b) ← infer_type h >>= get_rel_sides, + (a, b) ← infer_type h >>= instantiate_mvars >>= get_rel_sides, return $ (es.insert_list (get_nat_comps a)).insert_list (get_nat_comps b)) mk_rb_set, (++) l <$> nonnegs.to_list.mmap mk_coe_nat_nonneg_prf } @@ -183,7 +183,7 @@ into a proof of `t1 ≤ t2 + 1`. -/ meta def strengthen_strict_int : preprocessor := { name := "strengthen strict inequalities over int", transform := λ h, -do tp ← infer_type h, +do tp ← infer_type h >>= instantiate_mvars, guardb (is_strict_int_prop tp) >> singleton <$> mk_non_strict_int_pf_of_strict_int_pf h <|> return [h] } @@ -213,7 +213,7 @@ it tries to scale `t` to cancel out division by numerals. meta def cancel_denoms : preprocessor := { name := "cancel denominators", transform := λ pf, -(do some (_, lhs) ← parse_into_comp_and_expr <$> infer_type pf, +(do some (_, lhs) ← parse_into_comp_and_expr <$> (infer_type pf >>= instantiate_mvars), guardb $ lhs.contains_constant (= `has_div.div), singleton <$> normalize_denominators_in_lhs pf lhs) <|> return [pf] } @@ -272,7 +272,7 @@ This produces `2^n` branches when there are `n` such hypotheses in the input. -/ meta def remove_ne_aux : list expr → tactic (list branch) := λ hs, -(do e ← hs.mfind (λ e : expr, do e ← infer_type e, guard $ e.is_ne.is_some), +(do e ← hs.mfind (λ e : expr, do e ← infer_type e >>= instantiate_mvars, guard $ e.is_ne.is_some), [(_, ng1), (_, ng2)] ← to_expr ``(or.elim (lt_or_gt_of_ne %%e)) >>= apply, let do_goal : expr → tactic (list branch) := λ g, do set_goals [g], diff --git a/src/tactic/linarith/verification.lean b/src/tactic/linarith/verification.lean index 8f6e0b1e5010c..ea5123b0dd715 100644 --- a/src/tactic/linarith/verification.lean +++ b/src/tactic/linarith/verification.lean @@ -86,11 +86,11 @@ meta def mk_lt_zero_pf : list (expr × ℕ) → tactic expr /-- If `prf` is a proof of `t R s`, `term_of_ineq_prf prf` returns `t`. -/ meta def term_of_ineq_prf (prf : expr) : tactic expr := -prod.fst <$> (infer_type prf >>= get_rel_sides) +prod.fst <$> (infer_type prf >>= instantiate_mvars >>= get_rel_sides) /-- If `prf` is a proof of `t R s`, `ineq_prf_tp prf` returns the type of `t`. -/ meta def ineq_prf_tp (prf : expr) : tactic expr := -term_of_ineq_prf prf >>= infer_type +term_of_ineq_prf prf >>= infer_type >>= instantiate_mvars /-- `mk_neg_one_lt_zero_pf tp` returns a proof of `-1 < 0`, @@ -120,7 +120,7 @@ proof, it adds a proof of `-t = 0` to the list. meta def add_neg_eq_pfs : list expr → tactic (list expr) | [] := return [] | (h::t) := - do some (iq, tp) ← parse_into_comp_and_expr <$> infer_type h, + do some (iq, tp) ← parse_into_comp_and_expr <$> (infer_type h >>= instantiate_mvars), match iq with | ineq.eq := do nep ← mk_neg_eq_zero_pf h, tl ← add_neg_eq_pfs t, return $ h::nep::tl | _ := list.cons h <$> add_neg_eq_pfs t diff --git a/test/linarith.lean b/test/linarith.lean index 1d127062a60e5..365273aecad63 100644 --- a/test/linarith.lean +++ b/test/linarith.lean @@ -1,5 +1,14 @@ import tactic.linarith +example : ∀ (y : ℕ), y ≤ 37 → y < 40 := +begin + refine λ y hy, _, + -- The type of `hy` is a (solved but not instantiated) metavariable + do { tactic.get_local `hy >>= tactic.infer_type >>= guardb ∘ expr.is_mvar }, + -- But linarith should still work + linarith +end + example {α : Type} (_inst : Π (a : Prop), decidable a) [linear_ordered_field α] {a b c : α} From 88e6ad41d9c5d6da7e2ac45cae6d07a7fc6bb429 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sun, 16 Jul 2023 09:47:04 +0000 Subject: [PATCH 1258/1291] fix: handle archive and counterexamples correctly when adding port comments (#19237) --- scripts/add_port_comments.py | 7 +++++++ 1 file changed, 7 insertions(+) diff --git a/scripts/add_port_comments.py b/scripts/add_port_comments.py index 5c0160b5f9b56..366942aa8640f 100644 --- a/scripts/add_port_comments.py +++ b/scripts/add_port_comments.py @@ -12,6 +12,8 @@ status = PortStatus.deserialize_old() src_path = Path(__file__).parent.parent / 'src' +archive_path = Path(__file__).parent.parent / 'archive' +counterexamples_path = Path(__file__).parent.parent / 'counterexamples' def make_comment(fstatus): return textwrap.dedent(f"""\ @@ -87,6 +89,11 @@ def add_port_status(fcontent: str, fstatus: FileStatus) -> str: return fcontent def fname_for(import_path: str) -> Path: + for root in [src_path, archive_path, counterexamples_path]: + p = root / Path(*import_path.split('.')).with_suffix('.lean') + if p.exists(): + return p + # used only for error messages, a best-guess return src_path / Path(*import_path.split('.')).with_suffix('.lean') From 08b081ea92d80e3a41f899eea36ef6d56e0f1db0 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Sun, 16 Jul 2023 12:56:33 +0000 Subject: [PATCH 1259/1291] chore(*): add mathlib4 synchronization comments (#19238) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `algebra.expr` * `algebraic_geometry.morphisms.finite_type` * `algebraic_geometry.morphisms.ring_hom_properties` * `arithcc` * `canonically_ordered_comm_semiring_two_mul` * `char_p_zero_ne_char_zero` * `cyclotomic_105` * `data.matrix.auto` * `direct_sum_is_internal` * `examples.mersenne_primes` * `examples.prop_encodable` * `geometry.manifold.instances.sphere` * `girard` * `homogeneous_prime_not_prime` * `imo.imo1959_q1` * `imo.imo1960_q1` * `imo.imo1962_q1` * `imo.imo1962_q4` * `imo.imo1964_q1` * `imo.imo1969_q1` * `imo.imo1972_q5` * `imo.imo1975_q1` * `imo.imo1977_q6` * `imo.imo1981_q3` * `imo.imo1987_q1` * `imo.imo1988_q6` * `imo.imo1994_q1` * `imo.imo1998_q2` * `imo.imo2001_q2` * `imo.imo2001_q6` * `imo.imo2005_q3` * `imo.imo2005_q4` * `imo.imo2006_q3` * `imo.imo2006_q5` * `imo.imo2008_q2` * `imo.imo2008_q3` * `imo.imo2008_q4` * `imo.imo2011_q3` * `imo.imo2011_q5` * `imo.imo2013_q1` * `imo.imo2013_q5` * `imo.imo2019_q1` * `imo.imo2019_q2` * `imo.imo2019_q4` * `imo.imo2020_q2` * `imo.imo2021_q1` * `linear_order_with_pos_mul_pos_eq_zero` * `map_floor` * `miu_language.basic` * `miu_language.decision_nec` * `miu_language.decision_suf` * `oxford_invariants.2021summer.week3_p1` * `phillips` * `pseudoelement` * `quadratic_form` * `seminorm_lattice_not_distrib` * `sensitivity` * `sorgenfrey_line` * `wiedijk_100_theorems.abel_ruffini` * `wiedijk_100_theorems.area_of_a_circle` * `wiedijk_100_theorems.ascending_descending_sequences` * `wiedijk_100_theorems.ballot_problem` * `wiedijk_100_theorems.birthday_problem` * `wiedijk_100_theorems.cubing_a_cube` * `wiedijk_100_theorems.friendship_graphs` * `wiedijk_100_theorems.herons_formula` * `wiedijk_100_theorems.inverse_triangle_sum` * `wiedijk_100_theorems.konigsberg` * `wiedijk_100_theorems.partition` * `wiedijk_100_theorems.perfect_numbers` * `wiedijk_100_theorems.solution_of_cubic` * `wiedijk_100_theorems.sum_of_prime_reciprocals_diverges` * `zero_divisors_in_add_monoid_algebras` --- archive/arithcc.lean | 3 +++ archive/examples/mersenne_primes.lean | 3 +++ archive/examples/prop_encodable.lean | 3 +++ archive/imo/imo1959_q1.lean | 3 +++ archive/imo/imo1960_q1.lean | 3 +++ archive/imo/imo1962_q1.lean | 3 +++ archive/imo/imo1962_q4.lean | 3 +++ archive/imo/imo1964_q1.lean | 3 +++ archive/imo/imo1969_q1.lean | 3 +++ archive/imo/imo1972_q5.lean | 3 +++ archive/imo/imo1975_q1.lean | 3 +++ archive/imo/imo1977_q6.lean | 3 +++ archive/imo/imo1981_q3.lean | 3 +++ archive/imo/imo1987_q1.lean | 3 +++ archive/imo/imo1988_q6.lean | 3 +++ archive/imo/imo1994_q1.lean | 3 +++ archive/imo/imo1998_q2.lean | 3 +++ archive/imo/imo2001_q2.lean | 3 +++ archive/imo/imo2001_q6.lean | 3 +++ archive/imo/imo2005_q3.lean | 3 +++ archive/imo/imo2005_q4.lean | 3 +++ archive/imo/imo2006_q3.lean | 3 +++ archive/imo/imo2006_q5.lean | 3 +++ archive/imo/imo2008_q2.lean | 3 +++ archive/imo/imo2008_q3.lean | 3 +++ archive/imo/imo2008_q4.lean | 3 +++ archive/imo/imo2011_q3.lean | 3 +++ archive/imo/imo2011_q5.lean | 3 +++ archive/imo/imo2013_q1.lean | 3 +++ archive/imo/imo2013_q5.lean | 3 +++ archive/imo/imo2019_q1.lean | 3 +++ archive/imo/imo2019_q2.lean | 3 +++ archive/imo/imo2019_q4.lean | 3 +++ archive/imo/imo2020_q2.lean | 3 +++ archive/imo/imo2021_q1.lean | 3 +++ archive/miu_language/basic.lean | 3 +++ archive/miu_language/decision_nec.lean | 3 +++ archive/miu_language/decision_suf.lean | 3 +++ archive/oxford_invariants/2021summer/week3_p1.lean | 3 +++ archive/sensitivity.lean | 3 +++ archive/wiedijk_100_theorems/abel_ruffini.lean | 3 +++ archive/wiedijk_100_theorems/area_of_a_circle.lean | 3 +++ .../wiedijk_100_theorems/ascending_descending_sequences.lean | 3 +++ archive/wiedijk_100_theorems/ballot_problem.lean | 3 +++ archive/wiedijk_100_theorems/birthday_problem.lean | 3 +++ archive/wiedijk_100_theorems/cubing_a_cube.lean | 3 +++ archive/wiedijk_100_theorems/friendship_graphs.lean | 3 +++ archive/wiedijk_100_theorems/herons_formula.lean | 3 +++ archive/wiedijk_100_theorems/inverse_triangle_sum.lean | 3 +++ archive/wiedijk_100_theorems/konigsberg.lean | 3 +++ archive/wiedijk_100_theorems/partition.lean | 3 +++ archive/wiedijk_100_theorems/perfect_numbers.lean | 3 +++ archive/wiedijk_100_theorems/solution_of_cubic.lean | 3 +++ .../sum_of_prime_reciprocals_diverges.lean | 3 +++ counterexamples/canonically_ordered_comm_semiring_two_mul.lean | 3 +++ counterexamples/char_p_zero_ne_char_zero.lean | 3 +++ counterexamples/cyclotomic_105.lean | 3 +++ counterexamples/direct_sum_is_internal.lean | 3 +++ counterexamples/girard.lean | 3 +++ counterexamples/homogeneous_prime_not_prime.lean | 3 +++ counterexamples/linear_order_with_pos_mul_pos_eq_zero.lean | 3 +++ counterexamples/map_floor.lean | 3 +++ counterexamples/phillips.lean | 3 +++ counterexamples/pseudoelement.lean | 3 +++ counterexamples/quadratic_form.lean | 3 +++ counterexamples/seminorm_lattice_not_distrib.lean | 3 +++ counterexamples/sorgenfrey_line.lean | 3 +++ counterexamples/zero_divisors_in_add_monoid_algebras.lean | 3 +++ src/algebra/expr.lean | 3 +++ src/algebraic_geometry/morphisms/finite_type.lean | 3 +++ src/algebraic_geometry/morphisms/ring_hom_properties.lean | 3 +++ src/data/matrix/auto.lean | 3 +++ src/geometry/manifold/instances/sphere.lean | 3 +++ 73 files changed, 219 insertions(+) diff --git a/archive/arithcc.lean b/archive/arithcc.lean index 9de0770bbd8f2..f1feb07fa6173 100644 --- a/archive/arithcc.lean +++ b/archive/arithcc.lean @@ -9,6 +9,9 @@ import tactic.basic /-! # A compiler for arithmetic expressions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A formalization of the correctness of a compiler from arithmetic expressions to machine language described by McCarthy and Painter, which is considered the first proof of compiler correctness. diff --git a/archive/examples/mersenne_primes.lean b/archive/examples/mersenne_primes.lean index 2c5a28bdbea00..af4166ff31ebc 100644 --- a/archive/examples/mersenne_primes.lean +++ b/archive/examples/mersenne_primes.lean @@ -8,6 +8,9 @@ import number_theory.lucas_lehmer /-! # Explicit Mersenne primes +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We run some Lucas-Lehmer tests to prove some Mersenne primes are prime. See the discussion at the end of [src/number_theory/lucas_lehmer.lean] diff --git a/archive/examples/prop_encodable.lean b/archive/examples/prop_encodable.lean index a877eeed41334..d2436fd06f034 100644 --- a/archive/examples/prop_encodable.lean +++ b/archive/examples/prop_encodable.lean @@ -9,6 +9,9 @@ import data.W.basic /-! # W types +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The file `data/W.lean` shows that if `α` is an an encodable fintype and for every `a : α`, `β a` is encodable, then `W β` is encodable. diff --git a/archive/imo/imo1959_q1.lean b/archive/imo/imo1959_q1.lean index bb0c5f8b63b36..fe34d1b143489 100644 --- a/archive/imo/imo1959_q1.lean +++ b/archive/imo/imo1959_q1.lean @@ -10,6 +10,9 @@ import data.nat.prime /-! # IMO 1959 Q1 +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Prove that the fraction `(21n+4)/(14n+3)` is irreducible for every natural number `n`. Since Lean doesn't have a concept of "irreducible fractions" per se, we just formalize this diff --git a/archive/imo/imo1960_q1.lean b/archive/imo/imo1960_q1.lean index fa3ee7680092f..eccc3abb782e8 100644 --- a/archive/imo/imo1960_q1.lean +++ b/archive/imo/imo1960_q1.lean @@ -9,6 +9,9 @@ import data.nat.digits /-! # IMO 1960 Q1 +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Determine all three-digit numbers $N$ having the property that $N$ is divisible by 11, and $\dfrac{N}{11}$ is equal to the sum of the squares of the digits of $N$. diff --git a/archive/imo/imo1962_q1.lean b/archive/imo/imo1962_q1.lean index 0afc5e6be6726..bd586605533fb 100644 --- a/archive/imo/imo1962_q1.lean +++ b/archive/imo/imo1962_q1.lean @@ -9,6 +9,9 @@ import data.nat.digits /-! # IMO 1962 Q1 +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Find the smallest natural number $n$ which has the following properties: (a) Its decimal representation has 6 as the last digit. diff --git a/archive/imo/imo1962_q4.lean b/archive/imo/imo1962_q4.lean index 95b725d68beca..87996b4246960 100644 --- a/archive/imo/imo1962_q4.lean +++ b/archive/imo/imo1962_q4.lean @@ -9,6 +9,9 @@ import analysis.special_functions.trigonometric.complex /-! # IMO 1962 Q4 +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Solve the equation `cos x ^ 2 + cos (2 * x) ^ 2 + cos (3 * x) ^ 2 = 1`. Since Lean does not have a concept of "simplest form", we just express what is diff --git a/archive/imo/imo1964_q1.lean b/archive/imo/imo1964_q1.lean index fd987ba4f5f10..f381845902b31 100644 --- a/archive/imo/imo1964_q1.lean +++ b/archive/imo/imo1964_q1.lean @@ -10,6 +10,9 @@ import data.nat.modeq /-! # IMO 1964 Q1 +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + (a) Find all positive integers $n$ for which $2^n-1$ is divisible by $7$. (b) Prove that there is no positive integer $n$ for which $2^n+1$ is divisible by $7$. diff --git a/archive/imo/imo1969_q1.lean b/archive/imo/imo1969_q1.lean index 3eee264d353ca..b1e47a6bef7aa 100644 --- a/archive/imo/imo1969_q1.lean +++ b/archive/imo/imo1969_q1.lean @@ -11,6 +11,9 @@ import data.set.finite /-! # IMO 1969 Q1 +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Prove that there are infinitely many natural numbers $a$ with the following property: the number $z = n^4 + a$ is not prime for any natural number $n$. -/ diff --git a/archive/imo/imo1972_q5.lean b/archive/imo/imo1972_q5.lean index c796e589190f5..8dbb28ff56f04 100644 --- a/archive/imo/imo1972_q5.lean +++ b/archive/imo/imo1972_q5.lean @@ -10,6 +10,9 @@ import analysis.normed_space.basic /-! # IMO 1972 Q5 +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Problem: `f` and `g` are real-valued functions defined on the real line. For all `x` and `y`, `f(x + y) + f(x - y) = 2f(x)g(y)`. `f` is not identically zero and `|f(x)| ≤ 1` for all `x`. Prove that `|g(x)| ≤ 1` for all `x`. diff --git a/archive/imo/imo1975_q1.lean b/archive/imo/imo1975_q1.lean index eaf260b9cdbac..b135b1806671d 100644 --- a/archive/imo/imo1975_q1.lean +++ b/archive/imo/imo1975_q1.lean @@ -11,6 +11,9 @@ import algebra.big_operators.ring /-! # IMO 1975 Q1 +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Let `x₁, x₂, ... , xₙ` and `y₁, y₂, ... , yₙ` be two sequences of real numbers, such that `x₁ ≥ x₂ ≥ ... ≥ xₙ` and `y₁ ≥ y₂ ≥ ... ≥ yₙ`. Prove that if `z₁, z₂, ... , zₙ` is any permutation of `y₁, y₂, ... , yₙ`, then `∑ (xᵢ - yᵢ)^2 ≤ ∑ (xᵢ - zᵢ)^2` diff --git a/archive/imo/imo1977_q6.lean b/archive/imo/imo1977_q6.lean index 445644be7a341..8d3d44c050769 100644 --- a/archive/imo/imo1977_q6.lean +++ b/archive/imo/imo1977_q6.lean @@ -8,6 +8,9 @@ import data.pnat.basic /-! # IMO 1977 Q6 +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Suppose `f : ℕ+ → ℕ+` satisfies `f(f(n)) < f(n + 1)` for all `n`. Prove that `f(n) = n` for all `n`. diff --git a/archive/imo/imo1981_q3.lean b/archive/imo/imo1981_q3.lean index 4b40612e5a873..10a40331f2eec 100644 --- a/archive/imo/imo1981_q3.lean +++ b/archive/imo/imo1981_q3.lean @@ -10,6 +10,9 @@ import tactic.linarith /-! # IMO 1981 Q3 +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Determine the maximum value of `m ^ 2 + n ^ 2`, where `m` and `n` are integers in `{1, 2, ..., 1981}` and `(n ^ 2 - m * n - m ^ 2) ^ 2 = 1`. diff --git a/archive/imo/imo1987_q1.lean b/archive/imo/imo1987_q1.lean index b73054ddc9101..ffae4fbd202d7 100644 --- a/archive/imo/imo1987_q1.lean +++ b/archive/imo/imo1987_q1.lean @@ -11,6 +11,9 @@ import dynamics.fixed_points.basic /-! # Formalization of IMO 1987, Q1 +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Let $p_{n, k}$ be the number of permutations of a set of cardinality `n ≥ 1` that fix exactly `k` elements. Prove that $∑_{k=0}^n k p_{n,k}=n!$. diff --git a/archive/imo/imo1988_q6.lean b/archive/imo/imo1988_q6.lean index d5cd3fa2020e3..4e2699c0f61d9 100644 --- a/archive/imo/imo1988_q6.lean +++ b/archive/imo/imo1988_q6.lean @@ -13,6 +13,9 @@ import tactic.wlog /-! # IMO 1988 Q6 and constant descent Vieta jumping +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Question 6 of IMO1988 is somewhat (in)famous. Several expert problem solvers could not tackle the question within the given time limit. The problem lead to the introduction of a new proof technique, diff --git a/archive/imo/imo1994_q1.lean b/archive/imo/imo1994_q1.lean index 289fdf3250ad2..89b584ee3ebce 100644 --- a/archive/imo/imo1994_q1.lean +++ b/archive/imo/imo1994_q1.lean @@ -14,6 +14,9 @@ import tactic.by_contra /-! # IMO 1994 Q1 +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Let `m` and `n` be two positive integers. Let `a₁, a₂, ..., aₘ` be `m` different numbers from the set `{1, 2, ..., n}` such that for any two indices `i` and `j` with `1 ≤ i ≤ j ≤ m` and `aᵢ + aⱼ ≤ n`, diff --git a/archive/imo/imo1998_q2.lean b/archive/imo/imo1998_q2.lean index ab3cb7e1ccb35..7ce14e5cf3214 100644 --- a/archive/imo/imo1998_q2.lean +++ b/archive/imo/imo1998_q2.lean @@ -11,6 +11,9 @@ import tactic.noncomm_ring /-! # IMO 1998 Q2 + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. In a competition, there are `a` contestants and `b` judges, where `b ≥ 3` is an odd integer. Each judge rates each contestant as either "pass" or "fail". Suppose `k` is a number such that, for any two judges, their ratings coincide for at most `k` contestants. Prove that `k / a ≥ (b - 1) / (2b)`. diff --git a/archive/imo/imo2001_q2.lean b/archive/imo/imo2001_q2.lean index fc2d5172f39e8..346adc3db71fd 100644 --- a/archive/imo/imo2001_q2.lean +++ b/archive/imo/imo2001_q2.lean @@ -9,6 +9,9 @@ import analysis.special_functions.pow.real /-! # IMO 2001 Q2 +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Let $a$, $b$, $c$ be positive reals. Prove that $$ \frac{a}{\sqrt{a^2 + 8bc}} + diff --git a/archive/imo/imo2001_q6.lean b/archive/imo/imo2001_q6.lean index c2fa9f829a855..4b96c07090e1c 100644 --- a/archive/imo/imo2001_q6.lean +++ b/archive/imo/imo2001_q6.lean @@ -9,6 +9,9 @@ import tactic.linear_combination /-! # IMO 2001 Q6 + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. Let $a$, $b$, $c$, $d$ be integers with $a > b > c > d > 0$. Suppose that $$ a*c + b*d = (a + b - c + d) * (-a + b + c + d). $$ diff --git a/archive/imo/imo2005_q3.lean b/archive/imo/imo2005_q3.lean index bfcb287444ded..9d4ef279453ea 100644 --- a/archive/imo/imo2005_q3.lean +++ b/archive/imo/imo2005_q3.lean @@ -8,6 +8,9 @@ import tactic.positivity /-! # IMO 2005 Q3 + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. Let `x`, `y` and `z` be positive real numbers such that `xyz ≥ 1`. Prove that: `(x^5 - x^2)/(x^5 + y^2 + z^2) + (y^5 - y^2)/(y^5 + z^2 + x^2) + (z^5 - z^2)/(z^5 + x^2 + y^2) ≥ 0` diff --git a/archive/imo/imo2005_q4.lean b/archive/imo/imo2005_q4.lean index 9caceeddf5abf..ac90ef53f0bcc 100644 --- a/archive/imo/imo2005_q4.lean +++ b/archive/imo/imo2005_q4.lean @@ -8,6 +8,9 @@ import field_theory.finite.basic /-! # IMO 2005 Q4 +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Problem: Determine all positive integers relatively prime to all the terms of the infinite sequence `a n = 2 ^ n + 3 ^ n + 6 ^ n - 1`, for `n ≥ 1`. diff --git a/archive/imo/imo2006_q3.lean b/archive/imo/imo2006_q3.lean index 36a22e34bf2d0..8e3a0862d3ad4 100644 --- a/archive/imo/imo2006_q3.lean +++ b/archive/imo/imo2006_q3.lean @@ -9,6 +9,9 @@ import analysis.special_functions.sqrt /-! # IMO 2006 Q3 +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Determine the least real number $M$ such that $$ \left| ab(a^2 - b^2) + bc(b^2 - c^2) + ca(c^2 - a^2) \right| diff --git a/archive/imo/imo2006_q5.lean b/archive/imo/imo2006_q5.lean index 682b1994adb37..ec9929b2d4c1c 100644 --- a/archive/imo/imo2006_q5.lean +++ b/archive/imo/imo2006_q5.lean @@ -10,6 +10,9 @@ import dynamics.periodic_pts /-! # IMO 2006 Q5 +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Let $P(x)$ be a polynomial of degree $n>1$ with integer coefficients, and let $k$ be a positive integer. Consider the polynomial $Q(x) = P(P(\ldots P(P(x))\ldots))$, where $P$ occurs $k$ times. Prove that there are at most $n$ integers $t$ such that $Q(t)=t$. diff --git a/archive/imo/imo2008_q2.lean b/archive/imo/imo2008_q2.lean index eb4df1664207b..43c6b652cb127 100644 --- a/archive/imo/imo2008_q2.lean +++ b/archive/imo/imo2008_q2.lean @@ -8,6 +8,9 @@ import data.set.finite /-! # IMO 2008 Q2 + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. (a) Prove that ``` x^2 / (x-1)^2 + y^2 / (y-1)^2 + z^2 / (z-1)^2 ≥ 1 diff --git a/archive/imo/imo2008_q3.lean b/archive/imo/imo2008_q3.lean index 16a937f43fce5..ce1470e753e1b 100644 --- a/archive/imo/imo2008_q3.lean +++ b/archive/imo/imo2008_q3.lean @@ -12,6 +12,9 @@ import tactic.linear_combination /-! # IMO 2008 Q3 + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. Prove that there exist infinitely many positive integers `n` such that `n^2 + 1` has a prime divisor which is greater than `2n + √(2n)`. diff --git a/archive/imo/imo2008_q4.lean b/archive/imo/imo2008_q4.lean index 70339914ed2d3..c1cae17f559fa 100644 --- a/archive/imo/imo2008_q4.lean +++ b/archive/imo/imo2008_q4.lean @@ -10,6 +10,9 @@ import tactic.linear_combination /-! # IMO 2008 Q4 + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. Find all functions `f : (0,∞) → (0,∞)` (so, `f` is a function from the positive real numbers to the positive real numbers) such that ``` diff --git a/archive/imo/imo2011_q3.lean b/archive/imo/imo2011_q3.lean index c3e5dfc5d8361..67b4845b53018 100644 --- a/archive/imo/imo2011_q3.lean +++ b/archive/imo/imo2011_q3.lean @@ -9,6 +9,9 @@ import data.real.basic /-! # IMO 2011 Q3 +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Let f : ℝ → ℝ be a function that satisfies f(x + y) ≤ y * f(x) + f(f(x)) diff --git a/archive/imo/imo2011_q5.lean b/archive/imo/imo2011_q5.lean index abf6897ad352d..5bb51b977d729 100644 --- a/archive/imo/imo2011_q5.lean +++ b/archive/imo/imo2011_q5.lean @@ -8,6 +8,9 @@ import data.int.dvd.basic /-! # IMO 2011 Q5 +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Let `f` be a function from the set of integers to the set of positive integers. Suppose that, for any two integers `m` and `n`, the difference `f(m) - f(n)` is divisible by diff --git a/archive/imo/imo2013_q1.lean b/archive/imo/imo2013_q1.lean index cf7c03ce84b00..842d2de4c36f5 100644 --- a/archive/imo/imo2013_q1.lean +++ b/archive/imo/imo2013_q1.lean @@ -13,6 +13,9 @@ import tactic.field_simp /-! # IMO 2013 Q1 +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Prove that for any pair of positive integers k and n, there exist k positive integers m₁, m₂, ..., mₖ (not necessarily different) such that diff --git a/archive/imo/imo2013_q5.lean b/archive/imo/imo2013_q5.lean index d30c32da64014..16f7cdf6c6ab6 100644 --- a/archive/imo/imo2013_q5.lean +++ b/archive/imo/imo2013_q5.lean @@ -12,6 +12,9 @@ import tactic.positivity /-! # IMO 2013 Q5 +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Let `ℚ>₀` be the positive rational numbers. Let `f : ℚ>₀ → ℝ` be a function satisfying the conditions diff --git a/archive/imo/imo2019_q1.lean b/archive/imo/imo2019_q1.lean index 4e4772456a2ee..0226474656c48 100644 --- a/archive/imo/imo2019_q1.lean +++ b/archive/imo/imo2019_q1.lean @@ -8,6 +8,9 @@ import tactic.linarith /-! # IMO 2019 Q1 +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Determine all functions `f : ℤ → ℤ` such that, for all integers `a` and `b`, `f(2a) + 2f(b) = f(f(a+b))`. diff --git a/archive/imo/imo2019_q2.lean b/archive/imo/imo2019_q2.lean index 7be26dade77e2..b93b1046604ed 100644 --- a/archive/imo/imo2019_q2.lean +++ b/archive/imo/imo2019_q2.lean @@ -9,6 +9,9 @@ import geometry.euclidean.sphere.second_inter /-! # IMO 2019 Q2 +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In triangle `ABC`, point `A₁` lies on side `BC` and point `B₁` lies on side `AC`. Let `P` and `Q` be points on segments `AA₁` and `BB₁`, respectively, such that `PQ` is parallel to `AB`. Let `P₁` be a point on line `PB₁`, such that `B₁` lies strictly between `P` and `P₁`, and diff --git a/archive/imo/imo2019_q4.lean b/archive/imo/imo2019_q4.lean index 3465674493c60..05af48d4acb6e 100644 --- a/archive/imo/imo2019_q4.lean +++ b/archive/imo/imo2019_q4.lean @@ -10,6 +10,9 @@ import data.nat.multiplicity /-! # IMO 2019 Q4 +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Find all pairs `(k, n)` of positive integers such that ``` k! = (2 ^ n - 1)(2 ^ n - 2)(2 ^ n - 4)···(2 ^ n - 2 ^ (n - 1)) diff --git a/archive/imo/imo2020_q2.lean b/archive/imo/imo2020_q2.lean index beea58a78ff17..54938a0afb1ec 100644 --- a/archive/imo/imo2020_q2.lean +++ b/archive/imo/imo2020_q2.lean @@ -9,6 +9,9 @@ import analysis.mean_inequalities /-! # IMO 2020 Q2 +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The real numbers `a`, `b`, `c`, `d` are such that `a ≥ b ≥ c ≥ d > 0` and `a + b + c + d = 1`. Prove that `(a + 2b + 3c + 4d) a^a b^b c^c d^d < 1`. diff --git a/archive/imo/imo2021_q1.lean b/archive/imo/imo2021_q1.lean index 85abc0654bd36..397823fa856e1 100644 --- a/archive/imo/imo2021_q1.lean +++ b/archive/imo/imo2021_q1.lean @@ -14,6 +14,9 @@ import tactic.ring_exp /-! # IMO 2021 Q1 +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Let `n≥100` be an integer. Ivan writes the numbers `n, n+1,..., 2n` each on different cards. He then shuffles these `n+1` cards, and divides them into two piles. Prove that at least one of the piles contains two cards such that the sum of their numbers is a perfect square. diff --git a/archive/miu_language/basic.lean b/archive/miu_language/basic.lean index f4ca4329c312d..ad139d96ffe3b 100644 --- a/archive/miu_language/basic.lean +++ b/archive/miu_language/basic.lean @@ -8,6 +8,9 @@ import tactic.linarith /-! # An MIU Decision Procedure in Lean +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The [MIU formal system](https://en.wikipedia.org/wiki/MU_puzzle) was introduced by Douglas Hofstadter in the first chapter of his 1979 book, [Gödel, Escher, Bach](https://en.wikipedia.org/wiki/G%C3%B6del,_Escher,_Bach). diff --git a/archive/miu_language/decision_nec.lean b/archive/miu_language/decision_nec.lean index d70116a20b054..c6c5b287d1d28 100644 --- a/archive/miu_language/decision_nec.lean +++ b/archive/miu_language/decision_nec.lean @@ -11,6 +11,9 @@ import tactic.ring /-! # Decision procedure: necessary condition +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We introduce a condition `decstr` and show that if a string `en` is `derivable`, then `decstr en` holds. diff --git a/archive/miu_language/decision_suf.lean b/archive/miu_language/decision_suf.lean index 257bea77baa2e..8625a706de47b 100644 --- a/archive/miu_language/decision_suf.lean +++ b/archive/miu_language/decision_suf.lean @@ -9,6 +9,9 @@ import tactic.linarith /-! # Decision procedure - sufficient condition and decidability +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We give a sufficient condition for a string to be derivable in the MIU language. Together with the necessary condition, we use this to prove that `derivable` is an instance of `decidable_pred`. diff --git a/archive/oxford_invariants/2021summer/week3_p1.lean b/archive/oxford_invariants/2021summer/week3_p1.lean index eb30161c3a81c..cf1d4fe0e3ed0 100644 --- a/archive/oxford_invariants/2021summer/week3_p1.lean +++ b/archive/oxford_invariants/2021summer/week3_p1.lean @@ -11,6 +11,9 @@ import data.rat.cast /-! # The Oxford Invariants Puzzle Challenges - Summer 2021, Week 3, Problem 1 +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Original statement Let `n ≥ 3`, `a₁, ..., aₙ` be strictly positive integers such that `aᵢ ∣ aᵢ₋₁ + aᵢ₊₁` for diff --git a/archive/sensitivity.lean b/archive/sensitivity.lean index 995ff612ae9c3..1bd54993a14c4 100644 --- a/archive/sensitivity.lean +++ b/archive/sensitivity.lean @@ -15,6 +15,9 @@ import data.real.sqrt /-! # Huang's sensitivity theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A formalization of Hao Huang's sensitivity theorem: in the hypercube of dimension n ≥ 1, if one colors more than half the vertices then at least one vertex has at least √n colored neighbors. diff --git a/archive/wiedijk_100_theorems/abel_ruffini.lean b/archive/wiedijk_100_theorems/abel_ruffini.lean index 02b6bb4a7d250..16e3595a8dbf8 100644 --- a/archive/wiedijk_100_theorems/abel_ruffini.lean +++ b/archive/wiedijk_100_theorems/abel_ruffini.lean @@ -12,6 +12,9 @@ import ring_theory.eisenstein_criterion /-! # Construction of an algebraic number that is not solvable by radicals. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The main ingredients are: * `solvable_by_rad.is_solvable'` in `field_theory/abel_ruffini` : an irreducible polynomial with an `is_solvable_by_rad` root has solvable Galois group diff --git a/archive/wiedijk_100_theorems/area_of_a_circle.lean b/archive/wiedijk_100_theorems/area_of_a_circle.lean index e3c8d7a0daf6f..e22a6274e698f 100644 --- a/archive/wiedijk_100_theorems/area_of_a_circle.lean +++ b/archive/wiedijk_100_theorems/area_of_a_circle.lean @@ -11,6 +11,9 @@ import measure_theory.measure.lebesgue.integral /-! # Freek № 9: The Area of a Circle +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we show that the area of a disc with nonnegative radius `r` is `π * r^2`. The main tools our proof uses are `volume_region_between_eq_integral`, which allows us to represent the area of the disc as an integral, and `interval_integral.integral_eq_sub_of_has_deriv_at'_of_le`, the diff --git a/archive/wiedijk_100_theorems/ascending_descending_sequences.lean b/archive/wiedijk_100_theorems/ascending_descending_sequences.lean index a0d947a8089a5..14232f57243c4 100644 --- a/archive/wiedijk_100_theorems/ascending_descending_sequences.lean +++ b/archive/wiedijk_100_theorems/ascending_descending_sequences.lean @@ -8,6 +8,9 @@ import data.fintype.powerset /-! # Erdős–Szekeres theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves Theorem 73 from the [100 Theorems List](https://www.cs.ru.nl/~freek/100/), also known as the Erdős–Szekeres theorem: given a sequence of more than `r * s` distinct values, there is an increasing sequence of length longer than `r` or a decreasing sequence of length diff --git a/archive/wiedijk_100_theorems/ballot_problem.lean b/archive/wiedijk_100_theorems/ballot_problem.lean index a72c2e62f322e..f539b407b87e4 100644 --- a/archive/wiedijk_100_theorems/ballot_problem.lean +++ b/archive/wiedijk_100_theorems/ballot_problem.lean @@ -8,6 +8,9 @@ import probability.cond_count /-! # Ballot problem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves Theorem 30 from the [100 Theorems List](https://www.cs.ru.nl/~freek/100/). The ballot problem asks, if in an election, candidate A receives `p` votes whereas candidate B diff --git a/archive/wiedijk_100_theorems/birthday_problem.lean b/archive/wiedijk_100_theorems/birthday_problem.lean index d4544a61f7f70..4c0ee761bc19f 100644 --- a/archive/wiedijk_100_theorems/birthday_problem.lean +++ b/archive/wiedijk_100_theorems/birthday_problem.lean @@ -10,6 +10,9 @@ import probability.notation /-! # Birthday Problem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves Theorem 93 from the [100 Theorems List](https://www.cs.ru.nl/~freek/100/). As opposed to the standard probabilistic statement, we instead state the birthday problem diff --git a/archive/wiedijk_100_theorems/cubing_a_cube.lean b/archive/wiedijk_100_theorems/cubing_a_cube.lean index ab4b567df3ac4..f26ce9402eb3d 100644 --- a/archive/wiedijk_100_theorems/cubing_a_cube.lean +++ b/archive/wiedijk_100_theorems/cubing_a_cube.lean @@ -8,6 +8,9 @@ import data.set.finite import data.set.intervals.disjoint /-! +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Proof that a cube (in dimension n ≥ 3) cannot be cubed: There does not exist a partition of a cube into finitely many smaller cubes (at least two) of different sizes. diff --git a/archive/wiedijk_100_theorems/friendship_graphs.lean b/archive/wiedijk_100_theorems/friendship_graphs.lean index adf5f345655a9..222ca67a633ba 100644 --- a/archive/wiedijk_100_theorems/friendship_graphs.lean +++ b/archive/wiedijk_100_theorems/friendship_graphs.lean @@ -12,6 +12,9 @@ import tactic.interval_cases /-! # The Friendship Theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Definitions and Statement - A `friendship` graph is one in which any two distinct vertices have exactly one neighbor in common - A `politician`, at least in the context of this problem, is a vertex in a graph which is adjacent diff --git a/archive/wiedijk_100_theorems/herons_formula.lean b/archive/wiedijk_100_theorems/herons_formula.lean index 100102469b88c..0fb04262d0fe8 100644 --- a/archive/wiedijk_100_theorems/herons_formula.lean +++ b/archive/wiedijk_100_theorems/herons_formula.lean @@ -8,6 +8,9 @@ import geometry.euclidean.triangle /-! # Freek № 57: Heron's Formula +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves Theorem 57 from the [100 Theorems List](https://www.cs.ru.nl/~freek/100/), also known as Heron's formula, which gives the area of a triangle based only on its three sides' lengths. diff --git a/archive/wiedijk_100_theorems/inverse_triangle_sum.lean b/archive/wiedijk_100_theorems/inverse_triangle_sum.lean index 0b12999fbf3d8..b2a77cc1c4be9 100644 --- a/archive/wiedijk_100_theorems/inverse_triangle_sum.lean +++ b/archive/wiedijk_100_theorems/inverse_triangle_sum.lean @@ -9,6 +9,9 @@ import data.real.basic /-! # Sum of the Reciprocals of the Triangular Numbers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves Theorem 42 from the [100 Theorems List](https://www.cs.ru.nl/~freek/100/). We interpret “triangular numbers” as naturals of the form $\frac{k(k+1)}{2}$ for natural `k`. diff --git a/archive/wiedijk_100_theorems/konigsberg.lean b/archive/wiedijk_100_theorems/konigsberg.lean index 683d7314664db..fe5abdbe02a38 100644 --- a/archive/wiedijk_100_theorems/konigsberg.lean +++ b/archive/wiedijk_100_theorems/konigsberg.lean @@ -9,6 +9,9 @@ import tactic.derive_fintype /-! # The Königsberg bridges problem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We show that a graph that represents the islands and mainlands of Königsberg and seven bridges between them has no Eulerian trail. -/ diff --git a/archive/wiedijk_100_theorems/partition.lean b/archive/wiedijk_100_theorems/partition.lean index 8796e3b3c2a4d..ac3d5f6c41d05 100644 --- a/archive/wiedijk_100_theorems/partition.lean +++ b/archive/wiedijk_100_theorems/partition.lean @@ -15,6 +15,9 @@ import tactic.congrm /-! # Euler's Partition Theorem +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves Theorem 45 from the [100 Theorems List](https://www.cs.ru.nl/~freek/100/). The theorem concerns the counting of integer partitions -- ways of diff --git a/archive/wiedijk_100_theorems/perfect_numbers.lean b/archive/wiedijk_100_theorems/perfect_numbers.lean index 7d1cbd45655a7..191b9bd62c36e 100644 --- a/archive/wiedijk_100_theorems/perfect_numbers.lean +++ b/archive/wiedijk_100_theorems/perfect_numbers.lean @@ -12,6 +12,9 @@ import ring_theory.multiplicity /-! # Perfect Numbers +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves Theorem 70 from the [100 Theorems List](https://www.cs.ru.nl/~freek/100/). The theorem characterizes even perfect numbers. diff --git a/archive/wiedijk_100_theorems/solution_of_cubic.lean b/archive/wiedijk_100_theorems/solution_of_cubic.lean index 95b58db873e73..57e17a74e712f 100644 --- a/archive/wiedijk_100_theorems/solution_of_cubic.lean +++ b/archive/wiedijk_100_theorems/solution_of_cubic.lean @@ -9,6 +9,9 @@ import ring_theory.polynomial.cyclotomic.roots /-! # The Solution of a Cubic +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves Theorem 37 from the [100 Theorems List](https://www.cs.ru.nl/~freek/100/). In this file, we give the solutions to the cubic equation `a * x^3 + b * x^2 + c * x + d = 0` diff --git a/archive/wiedijk_100_theorems/sum_of_prime_reciprocals_diverges.lean b/archive/wiedijk_100_theorems/sum_of_prime_reciprocals_diverges.lean index 062c3d870d2d6..f52c458f2a87a 100644 --- a/archive/wiedijk_100_theorems/sum_of_prime_reciprocals_diverges.lean +++ b/archive/wiedijk_100_theorems/sum_of_prime_reciprocals_diverges.lean @@ -9,6 +9,9 @@ import data.nat.squarefree /-! # Divergence of the Prime Reciprocal Series +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file proves Theorem 81 from the [100 Theorems List](https://www.cs.ru.nl/~freek/100/). The theorem states that the sum of the reciprocals of all prime numbers diverges. The formalization follows Erdős's proof by upper and lower estimates. diff --git a/counterexamples/canonically_ordered_comm_semiring_two_mul.lean b/counterexamples/canonically_ordered_comm_semiring_two_mul.lean index 012d49af2dcee..624471d03aeb5 100644 --- a/counterexamples/canonically_ordered_comm_semiring_two_mul.lean +++ b/counterexamples/canonically_ordered_comm_semiring_two_mul.lean @@ -21,6 +21,9 @@ multiplication cannot be strengthened to **strict** monotonicity. Reference: https://leanprover.zulipchat.com/#narrow/stream/113489-new-members/topic/canonically_ordered.20pathology + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ namespace counterexample diff --git a/counterexamples/char_p_zero_ne_char_zero.lean b/counterexamples/char_p_zero_ne_char_zero.lean index a6b9cecde95cc..a567d155f59ec 100644 --- a/counterexamples/char_p_zero_ne_char_zero.lean +++ b/counterexamples/char_p_zero_ne_char_zero.lean @@ -7,6 +7,9 @@ import algebra.char_p.basic /-! # `char_p R 0` and `char_zero R` need not coincide for semirings +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + For rings, the two notions coincide. In fact, `char_p.of_char_zero` shows that `char_zero R` implies `char_p R 0` for any `char_zero` diff --git a/counterexamples/cyclotomic_105.lean b/counterexamples/cyclotomic_105.lean index cfbb5aa23bb94..9499e39f118ac 100644 --- a/counterexamples/cyclotomic_105.lean +++ b/counterexamples/cyclotomic_105.lean @@ -9,6 +9,9 @@ import ring_theory.polynomial.cyclotomic.basic /-! # Not all coefficients of cyclotomic polynomials are -1, 0, or 1 +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We show that not all coefficients of cyclotomic polynomials are equal to `0`, `-1` or `1`, in the theorem `not_forall_coeff_cyclotomic_neg_one_zero_one`. We prove this with the counterexample `coeff_cyclotomic_105 : coeff (cyclotomic 105 ℤ) 7 = -2`. diff --git a/counterexamples/direct_sum_is_internal.lean b/counterexamples/direct_sum_is_internal.lean index 9a4b75ece1807..08d293d5e7eb4 100644 --- a/counterexamples/direct_sum_is_internal.lean +++ b/counterexamples/direct_sum_is_internal.lean @@ -11,6 +11,9 @@ import tactic.fin_cases /-! # Not all complementary decompositions of a module over a semiring make up a direct sum +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This shows that while `ℤ≤0` and `ℤ≥0` are complementary `ℕ`-submodules of `ℤ`, which in turn implies as a collection they are `complete_lattice.independent` and that they span all of `ℤ`, they do not form a decomposition into a direct sum. diff --git a/counterexamples/girard.lean b/counterexamples/girard.lean index 090551161b38a..de764493e7027 100644 --- a/counterexamples/girard.lean +++ b/counterexamples/girard.lean @@ -8,6 +8,9 @@ import logic.basic /-! # Girard's paradox +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Girard's paradox is a proof that `Type : Type` entails a contradiction. We can't say this directly in Lean because `Type : Type 1` and it's not possible to give `Type` a different type via an axiom, so instead we axiomatize the behavior of the Pi type and application if the typing rule for Pi was diff --git a/counterexamples/homogeneous_prime_not_prime.lean b/counterexamples/homogeneous_prime_not_prime.lean index 1e895db644e53..41e0b98c88cdc 100644 --- a/counterexamples/homogeneous_prime_not_prime.lean +++ b/counterexamples/homogeneous_prime_not_prime.lean @@ -10,6 +10,9 @@ import tactic.derive_fintype /-! # A homogeneous prime that is homogeneously prime but not prime +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In `src/ring_theory/graded_algebra/radical.lean`, we assumed that the underline grading is indexed by a `linear_ordered_cancel_add_comm_monoid` to prove that a homogeneous ideal is prime if and only if it is homogeneously prime. This file is aimed to show that even if this assumption isn't strictly diff --git a/counterexamples/linear_order_with_pos_mul_pos_eq_zero.lean b/counterexamples/linear_order_with_pos_mul_pos_eq_zero.lean index 0fac502d4e5e3..b83f9dca5525b 100644 --- a/counterexamples/linear_order_with_pos_mul_pos_eq_zero.lean +++ b/counterexamples/linear_order_with_pos_mul_pos_eq_zero.lean @@ -16,6 +16,9 @@ The order is `0 < ε < 1`. Since `ε ^ 2 = 0`, the product of strictly positive Relevant Zulip chat: https://leanprover.zulipchat.com/#narrow/stream/116395-maths/topic/mul_pos + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ namespace counterexample diff --git a/counterexamples/map_floor.lean b/counterexamples/map_floor.lean index 2fd1b31c7cedc..105a24f6e9574 100644 --- a/counterexamples/map_floor.lean +++ b/counterexamples/map_floor.lean @@ -9,6 +9,9 @@ import data.polynomial.reverse /-! # Floors and ceils aren't preserved under ordered ring homomorphisms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Intuitively, if `f : α → β` is an ordered ring homomorphism, then floors and ceils should be preserved by `f` because: * `f` preserves the naturals/integers in `α` and `β` because it's a ring hom. diff --git a/counterexamples/phillips.lean b/counterexamples/phillips.lean index eafafd6d169bd..9ef4b12c2300f 100644 --- a/counterexamples/phillips.lean +++ b/counterexamples/phillips.lean @@ -11,6 +11,9 @@ import topology.continuous_function.bounded /-! # A counterexample on Pettis integrability +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + There are several theories of integration for functions taking values in Banach spaces. Bochner integration, requiring approximation by simple functions, is the analogue of the one-dimensional theory. It is very well behaved, but only works for functions with second-countable range. diff --git a/counterexamples/pseudoelement.lean b/counterexamples/pseudoelement.lean index 0e2ddaf8cc23b..a05f885f332ae 100644 --- a/counterexamples/pseudoelement.lean +++ b/counterexamples/pseudoelement.lean @@ -9,6 +9,9 @@ import algebra.category.Module.biproducts /-! # Pseudoelements and pullbacks + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. Borceux claims in Proposition 1.9.5 that the pseudoelement constructed in `category_theory.abelian.pseudoelement.pseudo_pullback` is unique. We show here that this claim is false. This means in particular that we cannot have an extensionality principle for pullbacks in diff --git a/counterexamples/quadratic_form.lean b/counterexamples/quadratic_form.lean index d9f029822dce1..50214d8f685dd 100644 --- a/counterexamples/quadratic_form.lean +++ b/counterexamples/quadratic_form.lean @@ -10,6 +10,9 @@ import data.zmod.basic /-! # `quadratic_form R M` and `subtype bilin_form.is_symm` are distinct notions in characteristic 2 +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The main result of this file is `bilin_form.not_inj_on_to_quadratic_form_is_symm`. The counterexample we use is $B (x, y) (x', y') ↦ xy' + x'y$ where `x y x' y' : zmod 2`. diff --git a/counterexamples/seminorm_lattice_not_distrib.lean b/counterexamples/seminorm_lattice_not_distrib.lean index 95b4e997ea397..ca9b74e8f6934 100644 --- a/counterexamples/seminorm_lattice_not_distrib.lean +++ b/counterexamples/seminorm_lattice_not_distrib.lean @@ -7,6 +7,9 @@ import analysis.seminorm /-! # The lattice of seminorms is not distributive +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We provide an example of three seminorms over the ℝ-vector space ℝ×ℝ which don't satisfy the lattice distributivity property `(p ⊔ q1) ⊓ (p ⊔ q2) ≤ p ⊔ (q1 ⊓ q2)`. diff --git a/counterexamples/sorgenfrey_line.lean b/counterexamples/sorgenfrey_line.lean index 065ecbecd57b5..ac740f1bb34da 100644 --- a/counterexamples/sorgenfrey_line.lean +++ b/counterexamples/sorgenfrey_line.lean @@ -13,6 +13,9 @@ import data.set.intervals.monotone /-! # Sorgenfrey line +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define `sorgenfrey_line` (notation: `ℝₗ`) to be the Sorgenfrey line. It is the real line with the topology space structure generated by half-open intervals `set.Ico a b`. diff --git a/counterexamples/zero_divisors_in_add_monoid_algebras.lean b/counterexamples/zero_divisors_in_add_monoid_algebras.lean index 88e430e411b30..8b52293f1958d 100644 --- a/counterexamples/zero_divisors_in_add_monoid_algebras.lean +++ b/counterexamples/zero_divisors_in_add_monoid_algebras.lean @@ -12,6 +12,9 @@ import data.zmod.basic /-! # Examples of zero-divisors in `add_monoid_algebra`s +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains an easy source of zero-divisors in an `add_monoid_algebra`. If `k` is a field and `G` is an additive group containing a non-zero torsion element, then `add_monoid_algebra k G` contains non-zero zero-divisors: this is lemma `zero_divisors_of_torsion`. diff --git a/src/algebra/expr.lean b/src/algebra/expr.lean index 7c6a862f0abfc..89641583e33bd 100644 --- a/src/algebra/expr.lean +++ b/src/algebra/expr.lean @@ -7,6 +7,9 @@ import tactic.core /-! ### Helpers to invoke functions involving algebra at tactic time +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + It's not clear whether using `instance_cache` is a sensible choice here. In particular, we need to use these tactics below when the algebraic instances are local variables that aren't in the "real" instance cache (the one used by `tactic.reset_instance_cache`). diff --git a/src/algebraic_geometry/morphisms/finite_type.lean b/src/algebraic_geometry/morphisms/finite_type.lean index 65ab36ddc83ec..ce417476ad834 100644 --- a/src/algebraic_geometry/morphisms/finite_type.lean +++ b/src/algebraic_geometry/morphisms/finite_type.lean @@ -9,6 +9,9 @@ import ring_theory.ring_hom.finite_type /-! # Morphisms of finite type +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + A morphism of schemes `f : X ⟶ Y` is locally of finite type if for each affine `U ⊆ Y` and `V ⊆ f ⁻¹' U`, The induced map `Γ(Y, U) ⟶ Γ(X, V)` is of finite type. diff --git a/src/algebraic_geometry/morphisms/ring_hom_properties.lean b/src/algebraic_geometry/morphisms/ring_hom_properties.lean index e389bdb008d2d..d30f2f44c77ec 100644 --- a/src/algebraic_geometry/morphisms/ring_hom_properties.lean +++ b/src/algebraic_geometry/morphisms/ring_hom_properties.lean @@ -10,6 +10,9 @@ import ring_theory.local_properties # Properties of morphisms from properties of ring homs. +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We provide the basic framework for talking about properties of morphisms that come from properties of ring homs. For `P` a property of ring homs, we have two ways of defining a property of scheme morphisms: diff --git a/src/data/matrix/auto.lean b/src/data/matrix/auto.lean index 1a351ef13828c..0d75a92737462 100644 --- a/src/data/matrix/auto.lean +++ b/src/data/matrix/auto.lean @@ -8,6 +8,9 @@ import data.matrix.reflection /-! # Automatically generated lemmas for working with concrete matrices +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file contains "magic" lemmas which autogenerate to the correct size of matrix. For instance, `matrix.of_mul_of_fin` can be used as: ```lean diff --git a/src/geometry/manifold/instances/sphere.lean b/src/geometry/manifold/instances/sphere.lean index e274632af5e28..71158b59e4aa9 100644 --- a/src/geometry/manifold/instances/sphere.lean +++ b/src/geometry/manifold/instances/sphere.lean @@ -15,6 +15,9 @@ import geometry.manifold.cont_mdiff_mfderiv /-! # Manifold structure on the sphere +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines stereographic projection from the sphere in an inner product space `E`, and uses it to put a smooth manifold structure on the sphere. From bf2428c9486c407ca38b5b3fb10b87dad0bc99fa Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sun, 16 Jul 2023 19:01:39 +0000 Subject: [PATCH 1260/1291] feat(order/irreducible): Sup-irreducible elements (#18999) Define sup- and inf- irreducible and prime elements in a lattice. --- src/data/fin/tuple/bubble_sort_induction.lean | 3 +- src/data/fintype/card.lean | 20 +- src/order/irreducible.lean | 229 ++++++++++++++++++ 3 files changed, 239 insertions(+), 13 deletions(-) create mode 100644 src/order/irreducible.lean diff --git a/src/data/fin/tuple/bubble_sort_induction.lean b/src/data/fin/tuple/bubble_sort_induction.lean index 8b496b9e8de09..8e7a0ca355e02 100644 --- a/src/data/fin/tuple/bubble_sort_induction.lean +++ b/src/data/fin/tuple/bubble_sort_induction.lean @@ -39,8 +39,7 @@ lemma bubble_sort_induction' {n : ℕ} {α : Type*} [linear_order α] {f : fin n P (f ∘ sort f) := begin letI := @preorder.lift _ (lex (fin n → α)) _ (λ σ : equiv.perm (fin n), to_lex (f ∘ σ)), - refine @well_founded.induction_bot' _ _ _ - (@finite.preorder.well_founded_lt (equiv.perm (fin n)) _ _) + refine @well_founded.induction_bot' _ _ _ (is_well_founded.wf : well_founded (<)) (equiv.refl _) (sort f) P (λ σ, f ∘ σ) (λ σ hσ hfσ, _) hf, obtain ⟨i, j, hij₁, hij₂⟩ := antitone_pair_of_not_sorted' hσ, exact ⟨σ * equiv.swap i j, pi.lex_desc hij₁ hij₂, h σ i j hij₁ hij₂ hfσ⟩, diff --git a/src/data/fintype/card.lean b/src/data/fintype/card.lean index 43d7bf3b7f0f5..be6e5f6be0c7c 100644 --- a/src/data/fintype/card.lean +++ b/src/data/fintype/card.lean @@ -726,17 +726,15 @@ have ∀ x y, r x y → (univ.filter (λ z, r z x)).card < (univ.filter (λ z, r exact ⟨λ z hzx, trans hzx hxy, not_forall_of_exists_not ⟨x, not_imp.2 ⟨hxy, irrefl x⟩⟩⟩, subrelation.wf this (measure_wf _) -lemma preorder.well_founded_lt [preorder α] : well_founded ((<) : α → α → Prop) := -well_founded_of_trans_of_irrefl _ - -lemma preorder.well_founded_gt [preorder α] : well_founded ((>) : α → α → Prop) := -well_founded_of_trans_of_irrefl _ - -@[priority 10] instance linear_order.is_well_order_lt [linear_order α] : is_well_order α (<) := -{ wf := preorder.well_founded_lt } - -@[priority 10] instance linear_order.is_well_order_gt [linear_order α] : is_well_order α (>) := -{ wf := preorder.well_founded_gt } +@[priority 100] -- See note [lower instance priority] +instance finite.to_well_founded_lt [preorder α] : well_founded_lt α := +⟨well_founded_of_trans_of_irrefl _⟩ +@[priority 100] -- See note [lower instance priority] +instance finite.to_well_founded_gt [preorder α] : well_founded_gt α := +⟨well_founded_of_trans_of_irrefl _⟩ + +@[priority 10] instance linear_order.is_well_order_lt [linear_order α] : is_well_order α (<) := {} +@[priority 10] instance linear_order.is_well_order_gt [linear_order α] : is_well_order α (>) := {} end finite diff --git a/src/order/irreducible.lean b/src/order/irreducible.lean new file mode 100644 index 0000000000000..07bec21729d7c --- /dev/null +++ b/src/order/irreducible.lean @@ -0,0 +1,229 @@ +/- +Copyright (c) 2023 Yaël Dillies. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies +-/ +import data.finset.lattice +import data.fintype.card + +/-! +# Irreducible and prime elements in an order + +This file defines irreducible and prime elements in an order and shows that in a well-founded +lattice every element decomposes as a supremum of irreducible elements. + +An element is sup-irreducible (resp. inf-irreducible) if it isn't `⊥` and can't be written as the +supremum of any strictly smaller elements. An element is sup-prime (resp. inf-prime) if it isn't `⊥` +and is greater than the supremum of any two elements less than it. + +Primality implies irreducibility in general. The converse only holds in distributive lattices. +Both hold for all (non-minimal) elements in a linear order. + +## Main declarations + +* `sup_irred a`: Sup-irreducibility, `a` isn't minimal and `a = b ⊔ c → a = b ∨ a = c` +* `inf_irred a`: Inf-irreducibility, `a` isn't maximal and `a = b ⊓ c → a = b ∨ a = c` +* `sup_prime a`: Sup-primality, `a` isn't minimal and `a ≤ b ⊔ c → a ≤ b ∨ a ≤ c` +* `inf_irred a`: Inf-primality, `a` isn't maximal and `a ≥ b ⊓ c → a ≥ b ∨ a ≥ c` +* `exists_sup_irred_decomposition`/`exists_inf_irred_decomposition`: Decomposition into irreducibles + in a well-founded semilattice. +-/ + +open finset order_dual + +variables {ι α : Type*} + +/-! ### Irreducible and prime elements -/ + +section semilattice_sup +variables [semilattice_sup α] {a b c : α} + +/-- A sup-irreducible element is a non-bottom element which isn't the supremum of anything smaller. +-/ +def sup_irred (a : α) : Prop := ¬ is_min a ∧ ∀ ⦃b c⦄, b ⊔ c = a → b = a ∨ c = a + +/-- A sup-prime element is a non-bottom element which isn't less than the supremum of anything +smaller. -/ +def sup_prime (a : α) : Prop := ¬ is_min a ∧ ∀ ⦃b c⦄, a ≤ b ⊔ c → a ≤ b ∨ a ≤ c + +lemma sup_irred.not_is_min (ha : sup_irred a) : ¬ is_min a := ha.1 +lemma sup_prime.not_is_min (ha : sup_prime a) : ¬ is_min a := ha.1 +lemma is_min.not_sup_irred (ha : is_min a) : ¬ sup_irred a := λ h, h.1 ha +lemma is_min.not_sup_prime (ha : is_min a) : ¬ sup_prime a := λ h, h.1 ha + +@[simp] lemma not_sup_irred : ¬ sup_irred a ↔ is_min a ∨ ∃ b c, b ⊔ c = a ∧ b < a ∧ c < a := +begin + rw [sup_irred, not_and_distrib], + push_neg, + rw exists₂_congr, + simp [@eq_comm _ _ a] { contextual := tt }, +end + +@[simp] lemma not_sup_prime : ¬ sup_prime a ↔ is_min a ∨ ∃ b c, a ≤ b ⊔ c ∧ ¬ a ≤ b ∧ ¬ a ≤ c := +by { rw [sup_prime, not_and_distrib], push_neg, refl } + +protected lemma sup_prime.sup_irred : sup_prime a → sup_irred a := +and.imp_right $ λ h b c ha, by simpa [←ha] using h ha.ge + +lemma sup_prime.le_sup (ha : sup_prime a) : a ≤ b ⊔ c ↔ a ≤ b ∨ a ≤ c := +⟨λ h, ha.2 h, λ h, h.elim le_sup_of_le_left le_sup_of_le_right⟩ + +variables [order_bot α] {s : finset ι} {f : ι → α} + +@[simp] lemma not_sup_irred_bot : ¬ sup_irred (⊥ : α) := is_min_bot.not_sup_irred +@[simp] lemma not_sup_prime_bot : ¬ sup_prime (⊥ : α) := is_min_bot.not_sup_prime + +lemma sup_irred.ne_bot (ha : sup_irred a) : a ≠ ⊥ := by { rintro rfl, exact not_sup_irred_bot ha } +lemma sup_prime.ne_bot (ha : sup_prime a) : a ≠ ⊥ := by { rintro rfl, exact not_sup_prime_bot ha } + +lemma sup_irred.finset_sup_eq (ha : sup_irred a) (h : s.sup f = a) : ∃ i ∈ s, f i = a := +begin + classical, + induction s using finset.induction with i s hi ih, + { simpa [ha.ne_bot] using h.symm }, + simp only [exists_prop, exists_mem_insert] at ⊢ ih, + rw sup_insert at h, + exact (ha.2 h).imp_right ih, +end + +lemma sup_prime.le_finset_sup (ha : sup_prime a) : a ≤ s.sup f ↔ ∃ i ∈ s, a ≤ f i := +begin + classical, + induction s using finset.induction with i s hi ih, + { simp [ha.ne_bot] }, + { simp only [exists_prop, exists_mem_insert, sup_insert, ha.le_sup, ih] } +end + +variables [well_founded_lt α] + +/-- In a well-founded lattice, any element is the supremum of finitely many sup-irreducible +elements. This is the order-theoretic analogue of prime factorisation. -/ +lemma exists_sup_irred_decomposition (a : α) : + ∃ s : finset α, s.sup id = a ∧ ∀ ⦃b⦄, b ∈ s → sup_irred b := +begin + classical, + apply well_founded_lt.induction a _, + clear a, + rintro a ih, + by_cases ha : sup_irred a, + { exact ⟨{a}, by simp [ha]⟩ }, + rw not_sup_irred at ha, + obtain ha | ⟨b, c, rfl, hb, hc⟩ := ha, + { exact ⟨∅, by simp [ha.eq_bot]⟩ }, + obtain ⟨s, rfl, hs⟩ := ih _ hb, + obtain ⟨t, rfl, ht⟩ := ih _ hc, + exact ⟨s ∪ t, sup_union, forall_mem_union.2 ⟨hs, ht⟩⟩, +end + +end semilattice_sup + +section semilattice_inf +variables [semilattice_inf α] {a b c : α} + +/-- An inf-irreducible element is a non-top element which isn't the infimum of anything bigger. -/ +def inf_irred (a : α) : Prop := ¬ is_max a ∧ ∀ ⦃b c⦄, b ⊓ c = a → b = a ∨ c = a + +/-- An inf-prime element is a non-top element which isn't bigger than the infimum of anything +bigger. -/ +def inf_prime (a : α) : Prop := ¬ is_max a ∧ ∀ ⦃b c⦄, b ⊓ c ≤ a → b ≤ a ∨ c ≤ a + +@[simp] lemma is_max.not_inf_irred (ha : is_max a) : ¬ inf_irred a := λ h, h.1 ha +@[simp] lemma is_max.not_inf_prime (ha : is_max a) : ¬ inf_prime a := λ h, h.1 ha + +@[simp] lemma not_inf_irred : ¬ inf_irred a ↔ is_max a ∨ ∃ b c, b ⊓ c = a ∧ a < b ∧ a < c := +@not_sup_irred αᵒᵈ _ _ + +@[simp] lemma not_inf_prime : ¬ inf_prime a ↔ is_max a ∨ ∃ b c, b ⊓ c ≤ a ∧ ¬ b ≤ a ∧ ¬ c ≤ a := +@not_sup_prime αᵒᵈ _ _ + +protected lemma inf_prime.inf_irred : inf_prime a → inf_irred a := +and.imp_right $ λ h b c ha, by simpa [←ha] using h ha.le + +lemma inf_prime.inf_le (ha : inf_prime a) : b ⊓ c ≤ a ↔ b ≤ a ∨ c ≤ a := +⟨λ h, ha.2 h, λ h, h.elim inf_le_of_left_le inf_le_of_right_le⟩ + +variables [order_top α] {s : finset ι} {f : ι → α} + +@[simp] lemma not_inf_irred_top : ¬ inf_irred (⊤ : α) := is_max_top.not_inf_irred +@[simp] lemma not_inf_prime_top : ¬ inf_prime (⊤ : α) := is_max_top.not_inf_prime + +lemma inf_irred.ne_top (ha : inf_irred a) : a ≠ ⊤ := by { rintro rfl, exact not_inf_irred_top ha } +lemma inf_prime.ne_top (ha : inf_prime a) : a ≠ ⊤ := by { rintro rfl, exact not_inf_prime_top ha } + +lemma inf_irred.finset_inf_eq : inf_irred a → s.inf f = a → ∃ i ∈ s, f i = a := +@sup_irred.finset_sup_eq _ αᵒᵈ _ _ _ _ _ + +lemma inf_prime.finset_inf_le (ha : inf_prime a) : s.inf f ≤ a ↔ ∃ i ∈ s, f i ≤ a := +@sup_prime.le_finset_sup _ αᵒᵈ _ _ _ _ _ ha + +variables [well_founded_gt α] + +/-- In a cowell-founded lattice, any element is the infimum of finitely many inf-irreducible +elements. This is the order-theoretic analogue of prime factorisation. -/ +lemma exists_inf_irred_decomposition (a : α) : + ∃ s : finset α, s.inf id = a ∧ ∀ ⦃b⦄, b ∈ s → inf_irred b := +@exists_sup_irred_decomposition αᵒᵈ _ _ _ _ + +end semilattice_inf + +section semilattice_sup +variables [semilattice_sup α] + +@[simp] lemma inf_irred_to_dual {a : α} : inf_irred (to_dual a) ↔ sup_irred a := iff.rfl +@[simp] lemma inf_prime_to_dual {a : α} : inf_prime (to_dual a) ↔ sup_prime a := iff.rfl +@[simp] lemma sup_irred_of_dual {a : αᵒᵈ} : sup_irred (of_dual a) ↔ inf_irred a := iff.rfl +@[simp] lemma sup_prime_of_dual {a : αᵒᵈ} : sup_prime (of_dual a) ↔ inf_prime a := iff.rfl + +alias inf_irred_to_dual ↔ _ sup_irred.dual +alias inf_prime_to_dual ↔ _ sup_prime.dual +alias sup_irred_of_dual ↔ _ inf_irred.of_dual +alias sup_prime_of_dual ↔ _ inf_prime.of_dual + +end semilattice_sup + +section semilattice_inf +variables [semilattice_inf α] + +@[simp] lemma sup_irred_to_dual {a : α} : sup_irred (to_dual a) ↔ inf_irred a := iff.rfl +@[simp] lemma sup_prime_to_dual {a : α} : sup_prime (to_dual a) ↔ inf_prime a := iff.rfl +@[simp] lemma inf_irred_of_dual {a : αᵒᵈ} : inf_irred (of_dual a) ↔ sup_irred a := iff.rfl +@[simp] lemma inf_prime_of_dual {a : αᵒᵈ} : inf_prime (of_dual a) ↔ sup_prime a := iff.rfl + +alias sup_irred_to_dual ↔ _ inf_irred.dual +alias sup_prime_to_dual ↔ _ inf_prime.dual +alias inf_irred_of_dual ↔ _ sup_irred.of_dual +alias inf_prime_of_dual ↔ _ sup_prime.of_dual + +end semilattice_inf + +section distrib_lattice +variables [distrib_lattice α] {a b c : α} + +@[simp] lemma sup_prime_iff_sup_irred : sup_prime a ↔ sup_irred a := +⟨sup_prime.sup_irred, and.imp_right $ λ h b c, + by { simp_rw [←inf_eq_left, inf_sup_left], exact @h _ _ }⟩ + +@[simp] lemma inf_prime_iff_inf_irred : inf_prime a ↔ inf_irred a := +⟨inf_prime.inf_irred, and.imp_right $ λ h b c, + by { simp_rw [←sup_eq_left, sup_inf_left], exact @h _ _ }⟩ + +alias sup_prime_iff_sup_irred ↔ _ sup_irred.sup_prime +alias inf_prime_iff_inf_irred ↔ _ inf_irred.inf_prime + +attribute [protected] sup_irred.sup_prime inf_irred.inf_prime + +end distrib_lattice + +section linear_order +variables [linear_order α] {a : α} + +@[simp] lemma sup_prime_iff_not_is_min : sup_prime a ↔ ¬ is_min a := and_iff_left $ by simp +@[simp] lemma inf_prime_iff_not_is_max : inf_prime a ↔ ¬ is_max a := and_iff_left $ by simp + +@[simp] lemma sup_irred_iff_not_is_min : sup_irred a ↔ ¬ is_min a := +and_iff_left $ λ _ _, by simpa only [sup_eq_max, max_eq_iff] using or.imp and.left and.left + +@[simp] lemma inf_irred_iff_not_is_max : inf_irred a ↔ ¬ is_max a := +and_iff_left $ λ _ _, by simpa only [inf_eq_min, min_eq_iff] using or.imp and.left and.left + +end linear_order From 8ea5598db6caeddde6cb734aa179cc2408dbd345 Mon Sep 17 00:00:00 2001 From: negiizhao Date: Wed, 19 Jul 2023 06:33:41 +0000 Subject: [PATCH 1261/1291] chore(set_theory/ordinal/basic): golf (#18547) Co-authored-by: Eric Wieser --- src/order/initial_seg.lean | 17 +++++++++ src/set_theory/ordinal/basic.lean | 60 ++++++++++--------------------- 2 files changed, 35 insertions(+), 42 deletions(-) diff --git a/src/order/initial_seg.lean b/src/order/initial_seg.lean index 2d3ef20adcc9d..e0365f4bc7537 100644 --- a/src/order/initial_seg.lean +++ b/src/order/initial_seg.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Floris van Doorn -/ +import logic.equiv.set import order.rel_iso.set import order.well_founded @@ -311,6 +312,22 @@ def of_element {α : Type*} (r : α → α → Prop) (a : α) : subrel r {b | r @[simp] theorem of_element_top {α : Type*} (r : α → α → Prop) (a : α) : (of_element r a).top = a := rfl +/-- For any principal segment `r ≺i s`, there is a `subrel` of `s` order isomorphic to `r`. -/ +@[simps symm_apply] +noncomputable def subrel_iso (f : r ≺i s) : subrel s {b | s b f.top} ≃r r := +rel_iso.symm +{ to_equiv := ((equiv.of_injective f f.injective).trans (equiv.set_congr + (funext (λ x, propext f.down.symm)))), + map_rel_iff' := λ a₁ a₂, f.map_rel_iff } + +@[simp] theorem apply_subrel_iso (f : r ≺i s) (b : {b | s b f.top}) : + f (f.subrel_iso b) = b := +equiv.apply_of_injective_symm f.injective _ + +@[simp] theorem subrel_iso_apply (f : r ≺i s) (a : α) : + f.subrel_iso ⟨f a, f.down.mpr ⟨a, rfl⟩⟩ = a := +equiv.of_injective_symm_apply f.injective _ + /-- Restrict the codomain of a principal segment -/ def cod_restrict (p : set β) (f : r ≺i s) (H : ∀ a, f a ∈ p) (H₂ : f.top ∈ p) : r ≺i subrel s p := diff --git a/src/set_theory/ordinal/basic.lean b/src/set_theory/ordinal/basic.lean index 2e146285a1f99..c9ccd28a73e48 100644 --- a/src/set_theory/ordinal/basic.lean +++ b/src/set_theory/ordinal/basic.lean @@ -358,28 +358,29 @@ injective_of_increasing r (<) (typein r) (λ x y, (typein_lt_typein r).2) {a b} : typein r a = typein r b ↔ a = b := (typein_injective r).eq_iff +/-- Principal segment version of the `typein` function, embedding a well order into + ordinals as a principal segment. -/ +def typein.principal_seg (r : α → α → Prop) [is_well_order α r] : + r ≺i ((<) : ordinal → ordinal → Prop) := +⟨⟨⟨typein r, typein_injective r⟩, λ a b, typein_lt_typein r⟩, type r, + λ o, ⟨typein_surj r, λ ⟨a, h⟩, h ▸ typein_lt_type r a⟩⟩ + +@[simp] theorem typein.principal_seg_coe (r : α → α → Prop) [is_well_order α r] : + (typein.principal_seg r : α → ordinal) = typein r := rfl + /-! ### Enumerating elements in a well-order with ordinals. -/ /-- `enum r o h` is the `o`-th element of `α` ordered by `r`. That is, `enum` maps an initial segment of the ordinals, those less than the order type of `r`, to the elements of `α`. -/ -def enum (r : α → α → Prop) [is_well_order α r] (o) : o < type r → α := -quot.rec_on o (λ ⟨β, s, _⟩ h, (classical.choice h).top) $ -λ ⟨β, s, _⟩ ⟨γ, t, _⟩ ⟨h⟩, begin - resetI, refine funext (λ (H₂ : type t < type r), _), - have H₁ : type s < type r, {rwa type_eq.2 ⟨h⟩}, - have : ∀ {o e} (H : o < type r), @@eq.rec - (λ (o : ordinal), o < type r → α) - (λ (h : type s < type r), (classical.choice h).top) - e H = (classical.choice H₁).top, {intros, subst e}, - exact (this H₂).trans (principal_seg.top_eq h - (classical.choice H₁) (classical.choice H₂)) -end +def enum (r : α → α → Prop) [is_well_order α r] (o) (h : o < type r) : α := +(typein.principal_seg r).subrel_iso ⟨o, h⟩ theorem enum_type {α β} {r : α → α → Prop} {s : β → β → Prop} [is_well_order α r] [is_well_order β s] (f : s ≺i r) {h : type s < type r} : enum r (type s) h = f.top := -principal_seg.top_eq (rel_iso.refl _) _ _ +(typein.principal_seg r).injective $ + ((typein.principal_seg r).apply_subrel_iso _).trans (typein_top _).symm @[simp] theorem enum_typein (r : α → α → Prop) [is_well_order α r] (a : α) : enum r (typein r a) (typein_lt_type r a) = a := @@ -412,13 +413,8 @@ lemma rel_iso_enum {α β : Type u} {r : α → α → Prop} {s : β → β → rel_iso_enum' _ _ _ _ theorem lt_wf : @well_founded ordinal (<) := -⟨λ a, induction_on a $ λ α r wo, by exactI -suffices ∀ a, acc (<) (typein r a), from -⟨_, λ o h, let ⟨a, e⟩ := typein_surj r h in e ▸ this a⟩, -λ a, acc.rec_on (wo.wf.apply a) $ λ x H IH, ⟨_, λ o h, begin - rcases typein_surj r (lt_trans h (typein_lt_type r _)) with ⟨b, rfl⟩, - exact IH _ ((typein_lt_typein r).1 h) -end⟩⟩ +well_founded_iff_well_founded_subrel.mpr $ λ a, induction_on a $ λ α r wo, by exactI + rel_hom_class.well_founded (typein.principal_seg r).subrel_iso wo.wf instance : has_well_founded ordinal := ⟨(<), lt_wf⟩ @@ -428,18 +424,6 @@ lemma induction {p : ordinal.{u} → Prop} (i : ordinal.{u}) (h : ∀ j, (∀ k, k < j → p k) → p j) : p i := lt_wf.induction i h -/-- Principal segment version of the `typein` function, embedding a well order into - ordinals as a principal segment. -/ -def typein.principal_seg {α : Type u} (r : α → α → Prop) [is_well_order α r] : - @principal_seg α ordinal.{u} r (<) := -⟨rel_embedding.of_monotone (typein r) - (λ a b, (typein_lt_typein r).2), type r, λ b, - ⟨λ h, ⟨enum r _ h, typein_enum r h⟩, - λ ⟨a, e⟩, e ▸ typein_lt_type _ _⟩⟩ - -@[simp] theorem typein.principal_seg_coe (r : α → α → Prop) [is_well_order α r] : - (typein.principal_seg r : α → ordinal) = typein r := rfl - /-! ### Cardinality of ordinals -/ /-- The cardinal of an ordinal is the cardinality of any type on which a relation with that order @@ -820,21 +804,13 @@ by { rw [←enum_typein (<) a, enum_le_enum', ←lt_succ_iff], apply typein_lt_s @[simp] theorem enum_inj {r : α → α → Prop} [is_well_order α r] {o₁ o₂ : ordinal} (h₁ : o₁ < type r) (h₂ : o₂ < type r) : enum r o₁ h₁ = enum r o₂ h₂ ↔ o₁ = o₂ := -⟨λ h, begin - by_contra hne, - cases lt_or_gt_of_ne hne with hlt hlt; - apply (is_well_order.is_irrefl r).1, - { rwa [←@enum_lt_enum α r _ o₁ o₂ h₁ h₂, h] at hlt }, - { change _ < _ at hlt, rwa [←@enum_lt_enum α r _ o₂ o₁ h₂ h₁, h] at hlt } -end, λ h, by simp_rw h⟩ +(typein.principal_seg r).subrel_iso.injective.eq_iff.trans subtype.mk_eq_mk /-- A well order `r` is order isomorphic to the set of ordinals smaller than `type r`. -/ @[simps] def enum_iso (r : α → α → Prop) [is_well_order α r] : subrel (<) (< type r) ≃r r := { to_fun := λ x, enum r x.1 x.2, inv_fun := λ x, ⟨typein r x, typein_lt_type r x⟩, - left_inv := λ ⟨o, h⟩, subtype.ext_val (typein_enum _ _), - right_inv := λ h, enum_typein _ _, - map_rel_iff' := by { rintros ⟨a, _⟩ ⟨b, _⟩, apply enum_lt_enum } } + ..(typein.principal_seg r).subrel_iso } /-- The order isomorphism between ordinals less than `o` and `o.out.α`. -/ @[simps] noncomputable def enum_iso_out (o : ordinal) : set.Iio o ≃o o.out.α := From 573eea921b01c49712ac02471911df0719297349 Mon Sep 17 00:00:00 2001 From: "github-actions[bot]" Date: Sat, 22 Jul 2023 09:12:42 +0000 Subject: [PATCH 1262/1291] chore(*): add mathlib4 synchronization comments (#19239) Regenerated from the [port status wiki page](https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status). Relates to the following files: * `order.irreducible` --- src/order/irreducible.lean | 3 +++ 1 file changed, 3 insertions(+) diff --git a/src/order/irreducible.lean b/src/order/irreducible.lean index 07bec21729d7c..0862574732a13 100644 --- a/src/order/irreducible.lean +++ b/src/order/irreducible.lean @@ -9,6 +9,9 @@ import data.fintype.card /-! # Irreducible and prime elements in an order +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines irreducible and prime elements in an order and shows that in a well-founded lattice every element decomposes as a supremum of irreducible elements. From 8047de4d911cdef39c2d646165eea972f7f9f539 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sat, 22 Jul 2023 18:41:27 +0000 Subject: [PATCH 1263/1291] feat(topology/metric_space/basic): decomposition of a "sphere" hypercube (#18875) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The main result here is `sphere x r = (⋃ i : β, function.eval i ⁻¹' sphere (x i) r) ∩ closed_ball x r`, which attempts to express that you can form the surface of a cube by taking the union of the faces in each axis. The `prod` result, `sphere (x, y) r = sphere x r ×ˢ closed_ball y r ∪ closed_ball x r ×ˢ sphere y r`, is a little easier to follow. I can imagine these being useful if we wanted to prove that the surface area of a "sphere" in `fin 3 -> R` was `24*r*r`! Co-authored-by: Yaël Dillies --- src/topology/metric_space/basic.lean | 66 +++++++++++++++++++++++++++- 1 file changed, 65 insertions(+), 1 deletion(-) diff --git a/src/topology/metric_space/basic.lean b/src/topology/metric_space/basic.lean index 4a34ea24186fa..de0baf9a61b0a 100644 --- a/src/topology/metric_space/basic.lean +++ b/src/topology/metric_space/basic.lean @@ -439,6 +439,12 @@ theorem mem_sphere' : y ∈ sphere x ε ↔ dist x y = ε := by rw [dist_comm, m theorem ne_of_mem_sphere (h : y ∈ sphere x ε) (hε : ε ≠ 0) : y ≠ x := by { contrapose! hε, symmetry, simpa [hε] using h } +theorem nonneg_of_mem_sphere (hy : y ∈ sphere x ε) : 0 ≤ ε := +dist_nonneg.trans_eq hy + +@[simp] theorem sphere_eq_empty_of_neg (hε : ε < 0) : sphere x ε = ∅ := +set.eq_empty_iff_forall_not_mem.mpr $ λ y hy, (nonneg_of_mem_sphere hy).not_lt hε + theorem sphere_eq_empty_of_subsingleton [subsingleton α] (hε : ε ≠ 0) : sphere x ε = ∅ := set.eq_empty_iff_forall_not_mem.mpr $ λ y hy, ne_of_mem_sphere hy hε (subsingleton.elim _ _) @@ -456,6 +462,10 @@ show dist x x ≤ ε, by rw dist_self; assumption @[simp] lemma closed_ball_eq_empty : closed_ball x ε = ∅ ↔ ε < 0 := by rw [← not_nonempty_iff_eq_empty, nonempty_closed_ball, not_le] +/-- Closed balls and spheres coincide when the radius is non-positive -/ +theorem closed_ball_eq_sphere_of_nonpos (hε : ε ≤ 0) : closed_ball x ε = sphere x ε := +set.ext $ λ _, (hε.trans dist_nonneg).le_iff_eq + theorem ball_subset_closed_ball : ball x ε ⊆ closed_ball x ε := assume y (hy : _ < _), le_of_lt hy @@ -1562,6 +1572,20 @@ theorem closed_ball_prod_same (x : α) (y : β) (r : ℝ) : closed_ball x r ×ˢ closed_ball y r = closed_ball (x, y) r := ext $ λ z, by simp [prod.dist_eq] +theorem sphere_prod (x : α × β) (r : ℝ) : + sphere x r = sphere x.1 r ×ˢ closed_ball x.2 r ∪ closed_ball x.1 r ×ˢ sphere x.2 r := +begin + obtain hr | rfl | hr := lt_trichotomy r 0, + { simp [hr], }, + { cases x, + simp_rw [←closed_ball_eq_sphere_of_nonpos le_rfl, union_self, closed_ball_prod_same] }, + { ext ⟨x', y'⟩, + simp_rw [set.mem_union, set.mem_prod, metric.mem_closed_ball, metric.mem_sphere, + prod.dist_eq, max_eq_iff], + refine or_congr (and_congr_right _) ((and_comm _ _).trans (and_congr_left _)), + all_goals { rintro rfl, refl } }, +end + end prod theorem uniform_continuous_dist : uniform_continuous (λp:α×α, dist p.1 p.2) := @@ -1802,11 +1826,25 @@ lemma nndist_pi_le_iff {f g : Πb, π b} {r : ℝ≥0} : nndist f g ≤ r ↔ ∀b, nndist (f b) (g b) ≤ r := by simp [nndist_pi_def] +lemma nndist_pi_lt_iff {f g : Πb, π b} {r : ℝ≥0} (hr : 0 < r) : + nndist f g < r ↔ ∀ b, nndist (f b) (g b) < r := +by simp [nndist_pi_def, finset.sup_lt_iff (show ⊥ < r, from hr)] + +lemma nndist_pi_eq_iff {f g : Π b, π b} {r : ℝ≥0} (hr : 0 < r) : + nndist f g = r ↔ (∃ i, nndist (f i) (g i) = r) ∧ ∀ b, nndist (f b) (g b) ≤ r := +begin + rw [eq_iff_le_not_lt, nndist_pi_lt_iff hr, nndist_pi_le_iff, not_forall, and_comm], + simp_rw [not_lt, and.congr_left_iff, le_antisymm_iff], + intro h, + refine exists_congr (λ b, _), + apply (and_iff_right $ h _).symm, +end + lemma dist_pi_lt_iff {f g : Πb, π b} {r : ℝ} (hr : 0 < r) : dist f g < r ↔ ∀b, dist (f b) (g b) < r := begin lift r to ℝ≥0 using hr.le, - simp [dist_pi_def, finset.sup_lt_iff (show ⊥ < r, from hr)], + exact nndist_pi_lt_iff hr, end lemma dist_pi_le_iff {f g : Πb, π b} {r : ℝ} (hr : 0 ≤ r) : @@ -1816,6 +1854,13 @@ begin exact nndist_pi_le_iff end +lemma dist_pi_eq_iff {f g : Πb, π b} {r : ℝ} (hr : 0 < r) : + dist f g = r ↔ (∃ i, dist (f i) (g i) = r) ∧ ∀ b, dist (f b) (g b) ≤ r := +begin + lift r to ℝ≥0 using hr.le, + simp_rw [←coe_nndist, nnreal.coe_eq, nndist_pi_eq_iff hr, nnreal.coe_le_coe], +end + lemma dist_pi_le_iff' [nonempty β] {f g : Π b, π b} {r : ℝ} : dist f g ≤ r ↔ ∀ b, dist (f b) (g b) ≤ r := begin @@ -1867,6 +1912,25 @@ lemma closed_ball_pi' [nonempty β] (x : Π b, π b) (r : ℝ) : closed_ball x r = set.pi univ (λ b, closed_ball (x b) r) := (le_or_lt 0 r).elim (closed_ball_pi x) $ λ hr, by simp [closed_ball_eq_empty.2 hr] +/-- A sphere in a product space is a union of spheres on each component restricted to the closed +ball. -/ +lemma sphere_pi (x : Πb, π b) {r : ℝ} (h : 0 < r ∨ nonempty β) : + sphere x r = (⋃ i : β, function.eval i ⁻¹' sphere (x i) r) ∩ closed_ball x r := +begin + obtain hr | rfl | hr := lt_trichotomy r 0, + { simp [hr], }, + { rw [closed_ball_eq_sphere_of_nonpos le_rfl, eq_comm, set.inter_eq_right_iff_subset], + letI := h.resolve_left (lt_irrefl _), + inhabit β, + refine subset_Union_of_subset default _, + intros x hx, + replace hx := hx.le, + rw [dist_pi_le_iff le_rfl] at hx, + exact le_antisymm (hx default) dist_nonneg }, + { ext, + simp [dist_pi_eq_iff hr, dist_pi_le_iff hr.le] }, +end + @[simp] lemma fin.nndist_insert_nth_insert_nth {n : ℕ} {α : fin (n + 1) → Type*} [Π i, pseudo_metric_space (α i)] (i : fin (n + 1)) (x y : α i) (f g : Π j, α (i.succ_above j)) : nndist (i.insert_nth x f) (i.insert_nth y g) = max (nndist x y) (nndist f g) := From 88fcdc3da43943f5b01925deddaa5bf0c0e85e4e Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sun, 23 Jul 2023 09:51:10 +0000 Subject: [PATCH 1264/1291] feat(ring_theory/tensor_product): add missing scalar tower instances (#19143) --- src/linear_algebra/tensor_product.lean | 11 ++++- src/ring_theory/tensor_product.lean | 65 ++++++++++++++++++++------ 2 files changed, 60 insertions(+), 16 deletions(-) diff --git a/src/linear_algebra/tensor_product.lean b/src/linear_algebra/tensor_product.lean index 7467afbf19bf2..2e031039e0618 100644 --- a/src/linear_algebra/tensor_product.lean +++ b/src/linear_algebra/tensor_product.lean @@ -293,7 +293,16 @@ section -- Like `R'`, `R'₂` provides a `distrib_mul_action R'₂ (M ⊗[R] N)` variables {R'₂ : Type*} [monoid R'₂] [distrib_mul_action R'₂ M] -variables [smul_comm_class R R'₂ M] [has_smul R'₂ R'] +variables [smul_comm_class R R'₂ M] + +/-- `smul_comm_class R' R'₂ M` implies `smul_comm_class R' R'₂ (M ⊗[R] N)` -/ +instance smul_comm_class_left [smul_comm_class R' R'₂ M] : smul_comm_class R' R'₂ (M ⊗[R] N) := +{ smul_comm := λ r' r'₂ x, tensor_product.induction_on x + (by simp_rw tensor_product.smul_zero) + (λ m n, by simp_rw [smul_tmul', smul_comm]) + (λ x y ihx ihy, by { simp_rw tensor_product.smul_add, rw [ihx, ihy] }),} + +variables [has_smul R'₂ R'] /-- `is_scalar_tower R'₂ R' M` implies `is_scalar_tower R'₂ R' (M ⊗[R] N)` -/ instance is_scalar_tower_left [is_scalar_tower R'₂ R' M] : diff --git a/src/ring_theory/tensor_product.lean b/src/ring_theory/tensor_product.lean index 9e3c0fdaa57d3..fe3a5edcf8001 100644 --- a/src/ring_theory/tensor_product.lean +++ b/src/ring_theory/tensor_product.lean @@ -322,16 +322,16 @@ begin { intros, simp only [linear_map.map_add, *, linear_map.add_apply], }, end -lemma mul_assoc (x y z : A ⊗[R] B) : mul (mul x y) z = mul x (mul y z) := +protected lemma mul_assoc (x y z : A ⊗[R] B) : mul (mul x y) z = mul x (mul y z) := mul_assoc' mul (by { intros, simp only [mul_apply, mul_assoc], }) x y z -lemma one_mul (x : A ⊗[R] B) : mul (1 ⊗ₜ 1) x = x := +protected lemma one_mul (x : A ⊗[R] B) : mul (1 ⊗ₜ 1) x = x := begin apply tensor_product.induction_on x; simp {contextual := tt}, end -lemma mul_one (x : A ⊗[R] B) : mul x (1 ⊗ₜ 1) = x := +protected lemma mul_one (x : A ⊗[R] B) : mul x (1 ⊗ₜ 1) = x := begin apply tensor_product.induction_on x; simp {contextual := tt}, @@ -347,9 +347,9 @@ instance : semiring (A ⊗[R] B) := add := (+), one := 1, mul := λ a b, mul a b, - one_mul := one_mul, - mul_one := mul_one, - mul_assoc := mul_assoc, + one_mul := algebra.tensor_product.one_mul, + mul_one := algebra.tensor_product.mul_one, + mul_assoc := algebra.tensor_product.mul_assoc, zero_mul := by simp, mul_zero := by simp, left_distrib := by simp, @@ -383,22 +383,57 @@ def include_left_ring_hom : A →+* A ⊗[R] B := map_one' := rfl, map_mul' := by simp } -variables {S : Type*} [comm_semiring S] [algebra S A] +variables {S : Type*} + +-- we want `is_scalar_tower_right` to take priority since it's better for unification elsewhere +@[priority 100] +instance is_scalar_tower_right + [monoid S] [distrib_mul_action S A] [is_scalar_tower S A A] [smul_comm_class R S A] : + is_scalar_tower S (A ⊗[R] B) (A ⊗[R] B) := +{ smul_assoc := λ r x y, begin + change (r • x) * y = r • (x * y), + apply tensor_product.induction_on y, + { simp [smul_zero], }, + { apply tensor_product.induction_on x, + { simp [smul_zero] }, + { intros a b a' b', + dsimp, + rw [tensor_product.smul_tmul', tensor_product.smul_tmul', tmul_mul_tmul, smul_mul_assoc], }, + { intros, simp [smul_add, add_mul, *], } }, + { intros, simp [smul_add, mul_add, *], }, + end } + +-- we want `algebra.to_smul_comm_class` to take priority since it's better for unification elsewhere +@[priority 100] +instance smul_comm_class_right + [monoid S] [distrib_mul_action S A] [smul_comm_class S A A] [smul_comm_class R S A] : + smul_comm_class S (A ⊗[R] B) (A ⊗[R] B) := +{ smul_comm := λ r x y, begin + change r • (x * y) = x * r • y, + apply tensor_product.induction_on y, + { simp [smul_zero], }, + { apply tensor_product.induction_on x, + { simp [smul_zero] }, + { intros a b a' b', + dsimp, + rw [tensor_product.smul_tmul', tensor_product.smul_tmul', tmul_mul_tmul, mul_smul_comm], }, + { intros, simp [smul_add, add_mul, *], } }, + { intros, simp [smul_add, mul_add, *], }, + end } + +variables [comm_semiring S] [algebra S A] instance left_algebra [smul_comm_class R S A] : algebra S (A ⊗[R] B) := { commutes' := λ r x, begin - apply tensor_product.induction_on x, - { simp, }, - { intros a b, dsimp, rw [algebra.commutes, _root_.mul_one, _root_.one_mul], }, - { intros y y' h h', dsimp at h h' ⊢, simp only [mul_add, add_mul, h, h'], }, + dsimp only [ring_hom.to_fun_eq_coe, ring_hom.comp_apply, include_left_ring_hom_apply], + rw [algebra_map_eq_smul_one, ←smul_tmul', ←one_def, mul_smul_comm, smul_mul_assoc, mul_one, + one_mul], end, smul_def' := λ r x, begin - apply tensor_product.induction_on x, - { simp [smul_zero], }, - { intros a b, dsimp, rw [tensor_product.smul_tmul', algebra.smul_def r a, _root_.one_mul] }, - { intros, dsimp, simp [smul_add, mul_add, *], }, + dsimp only [ring_hom.to_fun_eq_coe, ring_hom.comp_apply, include_left_ring_hom_apply], + rw [algebra_map_eq_smul_one, ←smul_tmul', smul_mul_assoc, ←one_def, one_mul], end, .. tensor_product.include_left_ring_hom.comp (algebra_map S A), .. (by apply_instance : module S (A ⊗[R] B)) }. From 4b92a463033b5587bb011657e25e4710bfca7364 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 24 Jul 2023 07:47:40 +0000 Subject: [PATCH 1265/1291] chore(ring_theory/kaehler): cleanup instances (#19234) The previous module instance has the wrong defeqs, and was more work to construct. The new one remains propositionally equal to the old one. --- src/ring_theory/ideal/cotangent.lean | 12 ++------ src/ring_theory/kaehler.lean | 45 +++++++++++----------------- 2 files changed, 21 insertions(+), 36 deletions(-) diff --git a/src/ring_theory/ideal/cotangent.lean b/src/ring_theory/ideal/cotangent.lean index 7743e96ac51d3..fe303936f6e4e 100644 --- a/src/ring_theory/ideal/cotangent.lean +++ b/src/ring_theory/ideal/cotangent.lean @@ -38,16 +38,10 @@ instance cotangent.module_of_tower : module S I.cotangent := submodule.quotient.module' _ instance : is_scalar_tower S S' I.cotangent := -begin - delta cotangent, - constructor, - intros s s' x, - rw [← @is_scalar_tower.algebra_map_smul S' R, ← @is_scalar_tower.algebra_map_smul S' R, - ← smul_assoc, ← is_scalar_tower.to_alg_hom_apply S S' R, map_smul], - refl -end +submodule.quotient.is_scalar_tower _ _ -instance [is_noetherian R I] : is_noetherian R I.cotangent := by { delta cotangent, apply_instance } +instance [is_noetherian R I] : is_noetherian R I.cotangent := +submodule.quotient.is_noetherian _ /-- The quotient map from `I` to `I ⧸ I ^ 2`. -/ @[simps apply (lemmas_only)] diff --git a/src/ring_theory/kaehler.lean b/src/ring_theory/kaehler.lean index ec3b0a3079dc8..2ec88c0e1d9fa 100644 --- a/src/ring_theory/kaehler.lean +++ b/src/ring_theory/kaehler.lean @@ -145,31 +145,23 @@ notation `Ω[`:100 S `⁄`:0 R `]`:0 := kaehler_differential R S instance : nonempty Ω[S⁄R] := ⟨0⟩ -instance kaehler_differential.module' {R' : Type*} [comm_ring R'] [algebra R' S] : +instance kaehler_differential.module' {R' : Type*} [comm_ring R'] [algebra R' S] + [smul_comm_class R R' S] : module R' Ω[S⁄R] := -(module.comp_hom (kaehler_differential.ideal R S).cotangent (algebra_map R' S) : _) +submodule.quotient.module' _ instance : is_scalar_tower S (S ⊗[R] S) Ω[S⁄R] := ideal.cotangent.is_scalar_tower _ instance kaehler_differential.is_scalar_tower_of_tower {R₁ R₂ : Type*} [comm_ring R₁] [comm_ring R₂] - [algebra R₁ S] [algebra R₂ S] [algebra R₁ R₂] [is_scalar_tower R₁ R₂ S] : + [algebra R₁ S] [algebra R₂ S] [has_smul R₁ R₂] + [smul_comm_class R R₁ S] [smul_comm_class R R₂ S] [is_scalar_tower R₁ R₂ S] : is_scalar_tower R₁ R₂ Ω[S⁄R] := -begin - convert restrict_scalars.is_scalar_tower R₁ R₂ Ω[S⁄R] using 1, - ext x m, - show algebra_map R₁ S x • m = algebra_map R₂ S (algebra_map R₁ R₂ x) • m, - rw ← is_scalar_tower.algebra_map_apply, -end +submodule.quotient.is_scalar_tower _ _ instance kaehler_differential.is_scalar_tower' : is_scalar_tower R (S ⊗[R] S) Ω[S⁄R] := -begin - convert restrict_scalars.is_scalar_tower R (S ⊗[R] S) Ω[S⁄R] using 1, - ext x m, - show algebra_map R S x • m = algebra_map R (S ⊗[R] S) x • m, - simp_rw [is_scalar_tower.algebra_map_apply R S (S ⊗[R] S), is_scalar_tower.algebra_map_smul] -end +submodule.quotient.is_scalar_tower _ _ /-- The quotient map `I → Ω[S⁄R]` with `I` being the kernel of `S ⊗[R] S → S`. -/ def kaehler_differential.from_ideal : kaehler_differential.ideal R S →ₗ[S ⊗[R] S] Ω[S⁄R] := @@ -190,13 +182,14 @@ rfl /-- The universal derivation into `Ω[S⁄R]`. -/ def kaehler_differential.D : derivation R S Ω[S⁄R] := { map_one_eq_zero' := begin - dsimp [kaehler_differential.D_linear_map_apply], + dsimp only [kaehler_differential.D_linear_map_apply], rw [ideal.to_cotangent_eq_zero, subtype.coe_mk, sub_self], exact zero_mem _ end, leibniz' := λ a b, begin - dsimp [kaehler_differential.D_linear_map_apply], - rw [← linear_map.map_smul_of_tower, ← linear_map.map_smul_of_tower, ← map_add, + dsimp only [kaehler_differential.D_linear_map_apply], + rw [← linear_map.map_smul_of_tower ((kaehler_differential.ideal R S).to_cotangent) a, + ← linear_map.map_smul_of_tower ((kaehler_differential.ideal R S).to_cotangent) b, ← map_add, ideal.to_cotangent_eq, pow_two], convert submodule.mul_mem_mul (kaehler_differential.one_smul_sub_smul_one_mem_ideal R a : _) (kaehler_differential.one_smul_sub_smul_one_mem_ideal R b : _) using 1, @@ -205,7 +198,7 @@ def kaehler_differential.D : derivation R S Ω[S⁄R] := submodule.coe_smul_of_tower, smul_sub, tensor_product.smul_tmul', smul_eq_mul, mul_one], ring_nf, end, - ..(kaehler_differential.D_linear_map R S) } + to_linear_map := kaehler_differential.D_linear_map R S } lemma kaehler_differential.D_apply (s : S) : kaehler_differential.D R S s = (kaehler_differential.ideal R S).to_cotangent @@ -373,9 +366,9 @@ into `(kaehler_differential.ideal R S).cotangent_ideal` -/ -- `derivation R S Ω[S⁄R] ≃ₗ[R] derivation R S (kaehler_differential.ideal R S).cotangent_ideal` -- But lean times-out if this is given explicitly. noncomputable -def kaehler_differential.End_equiv_derivation' := -@linear_equiv.comp_der R _ _ _ _ Ω[S⁄R] _ _ _ _ _ _ _ _ _ - ((kaehler_differential.ideal R S).cotangent_equiv_ideal.restrict_scalars S) +def kaehler_differential.End_equiv_derivation' : + derivation R S Ω[S⁄R] ≃ₗ[R] derivation R S _ := +linear_equiv.comp_der ((kaehler_differential.ideal R S).cotangent_equiv_ideal.restrict_scalars S) /-- (Implementation) An `equiv` version of `kaehler_differential.End_equiv_aux`. Used in `kaehler_differential.End_equiv`. -/ @@ -538,7 +531,7 @@ variables (A B : Type*) [comm_ring A] [comm_ring B] [algebra R A] [algebra S B] variables [algebra A B] [is_scalar_tower R S B] [is_scalar_tower R A B] -- The map `(A →₀ A) →ₗ[A] (B →₀ B)` -local notation `finsupp_map` := ((finsupp.map_range.linear_map (algebra.of_id A B).to_linear_map) +local notation `finsupp_map` := ((finsupp.map_range.linear_map (algebra.linear_map A B)) .comp (finsupp.lmap_domain A A (algebra_map A B))) lemma kaehler_differential.ker_total_map (h : function.surjective (algebra_map A B)) : @@ -552,7 +545,7 @@ begin set.image_univ, map_sub, map_add], simp only [linear_map.comp_apply, finsupp.map_range.linear_map_apply, finsupp.map_range_single, finsupp.lmap_domain_apply, finsupp.map_domain_single, alg_hom.to_linear_map_apply, - algebra.of_id_apply, ← is_scalar_tower.algebra_map_apply, map_one, map_add, map_mul], + algebra.linear_map_apply, ← is_scalar_tower.algebra_map_apply, map_one, map_add, map_mul], simp_rw [sup_assoc, ← (h.prod_map h).range_comp], congr' 3, rw [sup_eq_right], @@ -565,8 +558,6 @@ end presentation section exact_sequence -local attribute [irreducible] kaehler_differential - /- We have the commutative diagram A --→ B ↑ ↑ @@ -585,7 +576,7 @@ def derivation.comp_algebra_map [module A M] [module B M] [is_scalar_tower A B M leibniz' := λ a b, by simp, to_linear_map := d.to_linear_map.comp (is_scalar_tower.to_alg_hom R A B).to_linear_map } -variables (R B) +variables (R B) [smul_comm_class S A B] /-- The map `Ω[A⁄R] →ₗ[A] Ω[B⁄R]` given a square A --→ B From 148aefbd371a25f1cff33c85f20c661ce3155def Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 24 Jul 2023 08:57:01 +0000 Subject: [PATCH 1266/1291] docs(topology/dense_embedding): Improve explanation (#18134) An attempt at improving understandability. Closes #1539 --- src/topology/dense_embedding.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/topology/dense_embedding.lean b/src/topology/dense_embedding.lean index 3e768af19cd82..af0d11c6b63c5 100644 --- a/src/topology/dense_embedding.lean +++ b/src/topology/dense_embedding.lean @@ -15,8 +15,8 @@ import topology.bases This file defines three properties of functions: * `dense_range f` means `f` has dense image; -* `dense_inducing i` means `i` is also `inducing`; -* `dense_embedding e` means `e` is also an `embedding`. +* `dense_inducing i` means `i` is also `inducing`, namely it induces the topology on its codomain; +* `dense_embedding e` means `e` is further an `embedding`, namely it is injective and `inducing`. The main theorem `continuous_extend` gives a criterion for a function `f : X → Z` to a T₃ space Z to extend along a dense embedding From 4be589053caf347b899a494da75410deb55fb3ef Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 24 Jul 2023 08:57:03 +0000 Subject: [PATCH 1267/1291] fix(group_theory/subgroup/basic): generalize `centralizer` from `subgroup G` to `set G` (#18965) This is consistent with all the other `sub.centralizer` definitions. This generalization reveals that a lot of downstream results are rather strangely stated about `zpowers`. This does not attempt to change these, instead leaving the work for a follow up (either in a later mathlib3 PR or in mathlib4). --- src/group_theory/abelianization.lean | 11 +++++--- src/group_theory/commutator.lean | 2 +- src/group_theory/group_action/conj_act.lean | 3 +- src/group_theory/group_action/quotient.lean | 2 +- src/group_theory/subgroup/basic.lean | 31 +++++++++++---------- src/group_theory/subgroup/zpowers.lean | 10 ++++--- src/group_theory/sylow.lean | 12 ++++---- src/group_theory/transfer.lean | 2 +- 8 files changed, 41 insertions(+), 32 deletions(-) diff --git a/src/group_theory/abelianization.lean b/src/group_theory/abelianization.lean index 6c9fdf44175e5..c2915d7a36840 100644 --- a/src/group_theory/abelianization.lean +++ b/src/group_theory/abelianization.lean @@ -32,6 +32,8 @@ universes u v w -- Let G be a group. variables (G : Type u) [group G] +open subgroup (centralizer) + /-- The commutator subgroup of a group G is the normal subgroup generated by the commutators [p,q]=`p*q*p⁻¹*q⁻¹`. -/ @[derive subgroup.normal] @@ -64,12 +66,13 @@ begin end lemma commutator_centralizer_commutator_le_center : - ⁅(commutator G).centralizer, (commutator G).centralizer⁆ ≤ subgroup.center G := + ⁅centralizer (commutator G : set G), centralizer (commutator G : set G)⁆ ≤ subgroup.center G := begin - rw [←subgroup.centralizer_top, ←subgroup.commutator_eq_bot_iff_le_centralizer], - suffices : ⁅⁅⊤, (commutator G).centralizer⁆, (commutator G).centralizer⁆ = ⊥, + rw [←subgroup.centralizer_univ, ←subgroup.coe_top, + ←subgroup.commutator_eq_bot_iff_le_centralizer], + suffices : ⁅⁅⊤, centralizer (commutator G : set G)⁆, centralizer (commutator G : set G)⁆ = ⊥, { refine subgroup.commutator_commutator_eq_bot_of_rotate _ this, - rwa subgroup.commutator_comm (commutator G).centralizer }, + rwa subgroup.commutator_comm (centralizer (commutator G : set G)) }, rw [subgroup.commutator_comm, subgroup.commutator_eq_bot_iff_le_centralizer], exact set.centralizer_subset (subgroup.commutator_mono le_top le_top), end diff --git a/src/group_theory/commutator.lean b/src/group_theory/commutator.lean index 8d2ca4052460a..5c21d23e5f311 100644 --- a/src/group_theory/commutator.lean +++ b/src/group_theory/commutator.lean @@ -75,7 +75,7 @@ H₃.closure_le.trans ⟨λ h a b c d, h ⟨a, b, c, d, rfl⟩, λ h g ⟨a, b, lemma commutator_mono (h₁ : H₁ ≤ K₁) (h₂ : H₂ ≤ K₂) : ⁅H₁, H₂⁆ ≤ ⁅K₁, K₂⁆ := commutator_le.mpr (λ g₁ hg₁ g₂ hg₂, commutator_mem_commutator (h₁ hg₁) (h₂ hg₂)) -lemma commutator_eq_bot_iff_le_centralizer : ⁅H₁, H₂⁆ = ⊥ ↔ H₁ ≤ H₂.centralizer := +lemma commutator_eq_bot_iff_le_centralizer : ⁅H₁, H₂⁆ = ⊥ ↔ H₁ ≤ centralizer H₂ := begin rw [eq_bot_iff, commutator_le], refine forall_congr (λ p, forall_congr (λ hp, forall_congr (λ q, forall_congr (λ hq, _)))), diff --git a/src/group_theory/group_action/conj_act.lean b/src/group_theory/group_action/conj_act.lean index 12809e56e22d4..54cf1c4e85312 100644 --- a/src/group_theory/group_action/conj_act.lean +++ b/src/group_theory/group_action/conj_act.lean @@ -193,7 +193,8 @@ by { rw [is_conj_comm, is_conj_iff, mem_orbit_iff], refl } lemma orbit_rel_conj_act : (orbit_rel (conj_act G) G).rel = is_conj := funext₂ $ λ g h, by rw [orbit_rel_apply, mem_orbit_conj_act] -lemma stabilizer_eq_centralizer (g : G) : stabilizer (conj_act G) g = (zpowers g).centralizer := +lemma stabilizer_eq_centralizer (g : G) : + stabilizer (conj_act G) g = centralizer (zpowers (to_conj_act g) : set (conj_act G)) := le_antisymm (le_centralizer_iff.mp (zpowers_le.mpr (λ x, mul_inv_eq_iff_eq_mul.mp))) (λ x h, mul_inv_eq_of_eq_mul (h g (mem_zpowers g)).symm) diff --git a/src/group_theory/group_action/quotient.lean b/src/group_theory/group_action/quotient.lean index 0c2d4a21eb7af..d9bd002eff0e8 100644 --- a/src/group_theory/group_action/quotient.lean +++ b/src/group_theory/group_action/quotient.lean @@ -290,7 +290,7 @@ open quotient_group /-- Cosets of the centralizer of an element embed into the set of commutators. -/ noncomputable def quotient_centralizer_embedding (g : G) : - G ⧸ centralizer (zpowers (g : G)) ↪ commutator_set G := + G ⧸ centralizer (zpowers (g : G) : set G) ↪ commutator_set G := ((mul_action.orbit_equiv_quotient_stabilizer (conj_act G) g).trans (quotient_equiv_of_eq (conj_act.stabilizer_eq_centralizer g))).symm.to_embedding.trans ⟨λ x, ⟨x * g⁻¹, let ⟨_, x, rfl⟩ := x in ⟨x, g, rfl⟩⟩, λ x y, subtype.ext ∘ mul_right_cancel ∘ subtype.ext_iff.mp⟩ diff --git a/src/group_theory/subgroup/basic.lean b/src/group_theory/subgroup/basic.lean index 464a9504a5aaa..26cf2cc379965 100644 --- a/src/group_theory/subgroup/basic.lean +++ b/src/group_theory/subgroup/basic.lean @@ -1580,41 +1580,42 @@ normalizer_eq_top.mp (hmax.2 _ (hnc H (lt_top_iff_ne_top.mpr hmax.1))) end normalizer section centralizer +variables {H} /-- The `centralizer` of `H` is the subgroup of `g : G` commuting with every `h : H`. -/ @[to_additive "The `centralizer` of `H` is the additive subgroup of `g : G` commuting with every `h : H`."] -def centralizer : subgroup G := -{ carrier := set.centralizer H, +def centralizer (s : set G) : subgroup G := +{ carrier := set.centralizer s, inv_mem' := λ g, set.inv_mem_centralizer, - .. submonoid.centralizer ↑H } - -variables {H} + .. submonoid.centralizer s } -@[to_additive] lemma mem_centralizer_iff {g : G} : g ∈ H.centralizer ↔ ∀ h ∈ H, h * g = g * h := +@[to_additive] lemma mem_centralizer_iff {g : G} {s : set G} : + g ∈ centralizer s ↔ ∀ h ∈ s, h * g = g * h := iff.rfl -@[to_additive] lemma mem_centralizer_iff_commutator_eq_one {g : G} : - g ∈ H.centralizer ↔ ∀ h ∈ H, h * g * h⁻¹ * g⁻¹ = 1 := +@[to_additive] lemma mem_centralizer_iff_commutator_eq_one {g : G} {s : set G} : + g ∈ centralizer s ↔ ∀ h ∈ s, h * g * h⁻¹ * g⁻¹ = 1 := by simp only [mem_centralizer_iff, mul_inv_eq_iff_eq_mul, one_mul] -@[to_additive] lemma centralizer_top : centralizer ⊤ = center G := +@[to_additive] lemma centralizer_univ : centralizer set.univ = center G := set_like.ext' (set.centralizer_univ G) -@[to_additive] lemma le_centralizer_iff : H ≤ K.centralizer ↔ K ≤ H.centralizer := +@[to_additive] lemma le_centralizer_iff : H ≤ centralizer K ↔ K ≤ centralizer H := ⟨λ h x hx y hy, (h hy x hx).symm, λ h x hx y hy, (h hy x hx).symm⟩ @[to_additive] lemma center_le_centralizer (s) : center G ≤ centralizer s := set.center_subset_centralizer s -@[to_additive] lemma centralizer_le (h : H ≤ K) : centralizer K ≤ centralizer H := +@[to_additive] lemma centralizer_le {s t : set G} (h : s ⊆ t) : centralizer t ≤ centralizer s := submonoid.centralizer_le h -@[simp, to_additive] lemma centralizer_eq_top_iff_subset {s} : centralizer s = ⊤ ↔ s ≤ center G := +@[simp, to_additive] lemma centralizer_eq_top_iff_subset {s : set G} : + centralizer s = ⊤ ↔ s ⊆ center G := set_like.ext'_iff.trans set.centralizer_eq_top_iff_subset @[to_additive] instance subgroup.centralizer.characteristic [hH : H.characteristic] : - H.centralizer.characteristic := + (centralizer (H : set G)).characteristic := begin refine subgroup.characteristic_iff_comap_le.mpr (λ ϕ g hg h hh, ϕ.injective _), rw [map_mul, map_mul], @@ -1663,10 +1664,10 @@ end⟩⟩ (H.subgroup_of K).is_commutative := H.comap_injective_is_commutative subtype.coe_injective -@[to_additive] lemma le_centralizer_iff_is_commutative : K ≤ K.centralizer ↔ K.is_commutative := +@[to_additive] lemma le_centralizer_iff_is_commutative : K ≤ centralizer K ↔ K.is_commutative := ⟨λ h, ⟨⟨λ x y, subtype.ext (h y.2 x x.2)⟩⟩, λ h x hx y hy, congr_arg coe (h.1.1 ⟨y, hy⟩ ⟨x, hx⟩)⟩ -@[to_additive] lemma le_centralizer [h : H.is_commutative] : H ≤ H.centralizer := +@[to_additive] lemma le_centralizer [h : H.is_commutative] : H ≤ centralizer H := le_centralizer_iff_is_commutative.mpr h end subgroup diff --git a/src/group_theory/subgroup/zpowers.lean b/src/group_theory/subgroup/zpowers.lean index 6de48373ae8b7..f3456efdf70d3 100644 --- a/src/group_theory/subgroup/zpowers.lean +++ b/src/group_theory/subgroup/zpowers.lean @@ -166,17 +166,19 @@ zpowers_eq_bot.not subgroup.zpowers_eq_bot.mpr rfl @[to_additive] lemma centralizer_closure (S : set G) : - (closure S).centralizer = ⨅ g ∈ S, (zpowers g).centralizer := -le_antisymm (le_infi $ λ g, le_infi $ λ hg, centralizer_le $ zpowers_le.2 $ subset_closure hg) + centralizer (closure S : set G) = ⨅ g ∈ S, centralizer (zpowers g : set G) := +le_antisymm + (le_infi $ λ g, le_infi $ λ hg, centralizer_le $ set_like.coe_subset_coe.2 $ + zpowers_le.2 $ subset_closure hg) $ le_centralizer_iff.1 $ (closure_le _).2 $ λ g, set_like.mem_coe.2 ∘ zpowers_le.1 ∘ le_centralizer_iff.1 ∘ infi_le_of_le g ∘ infi_le _ @[to_additive] lemma center_eq_infi (S : set G) (hS : closure S = ⊤) : center G = ⨅ g ∈ S, centralizer (zpowers g) := -by rw [←centralizer_top, ←hS, centralizer_closure] +by rw [←centralizer_univ, ←coe_top, ←hS, centralizer_closure] @[to_additive] lemma center_eq_infi' (S : set G) (hS : closure S = ⊤) : - center G = ⨅ g : S, centralizer (zpowers g) := + center G = ⨅ g : S, centralizer (zpowers (g : G) : set G) := by rw [center_eq_infi S hS, ←infi_subtype''] end subgroup diff --git a/src/group_theory/sylow.lean b/src/group_theory/sylow.lean index 8b32b29b66805..1560211d6a95c 100644 --- a/src/group_theory/sylow.lean +++ b/src/group_theory/sylow.lean @@ -302,18 +302,19 @@ ext (λ g, sylow.smul_eq_iff_mem_normalizer) lemma sylow.conj_eq_normalizer_conj_of_mem_centralizer [fact p.prime] [finite (sylow p G)] (P : sylow p G) (x g : G) - (hx : x ∈ (P : subgroup G).centralizer) (hy : g⁻¹ * x * g ∈ (P : subgroup G).centralizer) : + (hx : x ∈ centralizer (P : set G)) (hy : g⁻¹ * x * g ∈ centralizer (P : set G)) : ∃ n ∈ (P : subgroup G).normalizer, g⁻¹ * x * g = n⁻¹ * x * n := begin - have h1 : ↑P ≤ (zpowers x).centralizer, + have h1 : ↑P ≤ centralizer (zpowers x : set G), { rwa [le_centralizer_iff, zpowers_le] }, - have h2 : ↑(g • P) ≤ (zpowers x).centralizer, + have h2 : ↑(g • P) ≤ centralizer (zpowers x : set G), { rw [le_centralizer_iff, zpowers_le], rintros - ⟨z, hz, rfl⟩, specialize hy z hz, rwa [←mul_assoc, ←eq_mul_inv_iff_mul_eq, mul_assoc, mul_assoc, mul_assoc, ←mul_assoc, eq_inv_mul_iff_mul_eq, ←mul_assoc, ←mul_assoc] at hy }, - obtain ⟨h, hh⟩ := exists_smul_eq (zpowers x).centralizer ((g • P).subtype h2) (P.subtype h1), + obtain ⟨h, hh⟩ := + exists_smul_eq (centralizer (zpowers x : set G)) ((g • P).subtype h2) (P.subtype h1), simp_rw [sylow.smul_subtype, smul_def, smul_smul] at hh, refine ⟨h * g, sylow.smul_eq_iff_mem_normalizer.mp (sylow.subtype_injective hh), _⟩, rw [←mul_assoc, commute.right_comm (h.prop x (mem_zpowers x)), mul_inv_rev, inv_mul_cancel_right] @@ -322,7 +323,8 @@ end lemma sylow.conj_eq_normalizer_conj_of_mem [fact p.prime] [finite (sylow p G)] (P : sylow p G) [hP : (P : subgroup G).is_commutative] (x g : G) (hx : x ∈ P) (hy : g⁻¹ * x * g ∈ P) : ∃ n ∈ (P : subgroup G).normalizer, g⁻¹ * x * g = n⁻¹ * x * n := -P.conj_eq_normalizer_conj_of_mem_centralizer x g (le_centralizer P hx) (le_centralizer P hy) +P.conj_eq_normalizer_conj_of_mem_centralizer x g + (le_centralizer (P : subgroup G) hx : _) (le_centralizer (P : subgroup G) hy : _) /-- Sylow `p`-subgroups are in bijection with cosets of the normalizer of a Sylow `p`-subgroup -/ noncomputable def sylow.equiv_quotient_normalizer [fact p.prime] [fintype (sylow p G)] diff --git a/src/group_theory/transfer.lean b/src/group_theory/transfer.lean index f54f4a3cc4f1d..1202268a0a7b7 100644 --- a/src/group_theory/transfer.lean +++ b/src/group_theory/transfer.lean @@ -172,7 +172,7 @@ rfl section burnside_transfer -variables {p : ℕ} (P : sylow p G) (hP : (P : subgroup G).normalizer ≤ (P : subgroup G).centralizer) +variables {p : ℕ} (P : sylow p G) (hP : (P : subgroup G).normalizer ≤ centralizer (P : set G)) include hP From c0c52abb75074ed8b73a948341f50521fbf43b4c Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 24 Jul 2023 11:50:08 +0000 Subject: [PATCH 1268/1291] feat(order/upper_lower/basic): Linear order (#19068) Upper/lower sets on a linear order themselves form a linear order. --- src/algebra/order/upper_lower.lean | 15 +++--- src/order/upper_lower/basic.lean | 53 +++++++++++++++++++-- src/topology/algebra/order/upper_lower.lean | 2 +- 3 files changed, 56 insertions(+), 14 deletions(-) diff --git a/src/algebra/order/upper_lower.lean b/src/algebra/order/upper_lower.lean index 3f36d91f00fa1..d52a2b05f96bc 100644 --- a/src/algebra/order/upper_lower.lean +++ b/src/algebra/order/upper_lower.lean @@ -34,13 +34,10 @@ section ordered_comm_group variables {α : Type*} [ordered_comm_group α] {s t : set α} {a : α} @[to_additive] lemma is_upper_set.smul (hs : is_upper_set s) : is_upper_set (a • s) := -begin - rintro _ y hxy ⟨x, hx, rfl⟩, - exact mem_smul_set_iff_inv_smul_mem.2 (hs (le_inv_mul_iff_mul_le.2 hxy) hx), -end +hs.image $ order_iso.mul_left _ @[to_additive] lemma is_lower_set.smul (hs : is_lower_set s) : is_lower_set (a • s) := -hs.of_dual.smul +hs.image $ order_iso.mul_left _ @[to_additive] lemma set.ord_connected.smul (hs : s.ord_connected) : (a • s).ord_connected := begin @@ -55,10 +52,10 @@ by { rw [←smul_eq_mul, ←bUnion_smul_set], exact is_upper_set_Union₂ (λ x by { rw mul_comm, exact hs.mul_left } @[to_additive] lemma is_lower_set.mul_left (ht : is_lower_set t) : is_lower_set (s * t) := -ht.of_dual.mul_left +ht.to_dual.mul_left @[to_additive] lemma is_lower_set.mul_right (hs : is_lower_set s) : is_lower_set (s * t) := -hs.of_dual.mul_right +hs.to_dual.mul_right @[to_additive] lemma is_upper_set.inv (hs : is_upper_set s) : is_lower_set s⁻¹ := λ x y h, hs $ inv_le_inv' h @@ -73,10 +70,10 @@ by { rw div_eq_mul_inv, exact ht.inv.mul_left } by { rw div_eq_mul_inv, exact hs.mul_right } @[to_additive] lemma is_lower_set.div_left (ht : is_lower_set t) : is_upper_set (s / t) := -ht.of_dual.div_left +ht.to_dual.div_left @[to_additive] lemma is_lower_set.div_right (hs : is_lower_set s) : is_lower_set (s / t) := -hs.of_dual.div_right +hs.to_dual.div_right namespace upper_set diff --git a/src/order/upper_lower/basic.lean b/src/order/upper_lower/basic.lean index d95fe66e45e02..db5b4c91df203 100644 --- a/src/order/upper_lower/basic.lean +++ b/src/order/upper_lower/basic.lean @@ -6,6 +6,7 @@ Authors: Yaël Dillies, Sara Rousta import data.set_like.basic import data.set.intervals.ord_connected import data.set.intervals.order_iso +import tactic.by_contra /-! # Up-sets and down-sets @@ -135,10 +136,10 @@ iff.rfl @[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 -alias is_upper_set_preimage_of_dual_iff ↔ _ is_lower_set.of_dual -alias is_lower_set_preimage_to_dual_iff ↔ _ is_upper_set.to_dual -alias is_upper_set_preimage_to_dual_iff ↔ _ is_lower_set.to_dual +alias is_lower_set_preimage_of_dual_iff ↔ _ is_upper_set.to_dual +alias is_upper_set_preimage_of_dual_iff ↔ _ is_lower_set.to_dual +alias is_lower_set_preimage_to_dual_iff ↔ _ is_upper_set.of_dual +alias is_upper_set_preimage_to_dual_iff ↔ _ is_lower_set.of_dual end has_le @@ -266,6 +267,24 @@ alias is_lower_set_iff_Iio_subset ↔ is_lower_set.Iio_subset _ end partial_order +section linear_order +variables [linear_order α] {s t : set α} + +lemma is_upper_set.total (hs : is_upper_set s) (ht : is_upper_set t) : s ⊆ t ∨ t ⊆ s := +begin + by_contra' h, + simp_rw set.not_subset at h, + obtain ⟨⟨a, has, hat⟩, b, hbt, hbs⟩ := h, + obtain hab | hba := le_total a b, + { exact hbs (hs hab has) }, + { exact hat (ht hba hbt) } +end + +lemma is_lower_set.total (hs : is_lower_set s) (ht : is_lower_set t) : s ⊆ t ∨ t ⊆ s := +hs.to_dual.total ht.to_dual + +end linear_order + /-! ### Bundled upper/lower sets -/ section has_le @@ -519,6 +538,32 @@ end lower_set end has_le +section linear_order +variables [linear_order α] + +instance upper_set.is_total_le : is_total (upper_set α) (≤) := ⟨λ s t, t.upper.total s.upper⟩ +instance lower_set.is_total_le : is_total (lower_set α) (≤) := ⟨λ s t, s.lower.total t.lower⟩ + +noncomputable instance : complete_linear_order (upper_set α) := +{ le_total := is_total.total, + decidable_le := classical.dec_rel _, + decidable_eq := classical.dec_rel _, + decidable_lt := classical.dec_rel _, + max_def := by classical; exact sup_eq_max_default, + min_def := by classical; exact inf_eq_min_default, + ..upper_set.complete_distrib_lattice } + +noncomputable instance : complete_linear_order (lower_set α) := +{ le_total := is_total.total, + decidable_le := classical.dec_rel _, + decidable_eq := classical.dec_rel _, + decidable_lt := classical.dec_rel _, + max_def := by classical; exact sup_eq_max_default, + min_def := by classical; exact inf_eq_min_default, + ..lower_set.complete_distrib_lattice } + +end linear_order + /-! #### Map -/ section diff --git a/src/topology/algebra/order/upper_lower.lean b/src/topology/algebra/order/upper_lower.lean index 5bc75c68c9be3..792ad0b431a15 100644 --- a/src/topology/algebra/order/upper_lower.lean +++ b/src/topology/algebra/order/upper_lower.lean @@ -85,7 +85,7 @@ protected lemma is_upper_set.interior (h : is_upper_set s) : is_upper_set (inter by { rw [←is_lower_set_compl, ←closure_compl], exact h.compl.closure } protected lemma is_lower_set.interior (h : is_lower_set s) : is_lower_set (interior s) := -h.of_dual.interior +h.to_dual.interior protected lemma set.ord_connected.interior (h : s.ord_connected) : (interior s).ord_connected := begin From 18ee599842a5d17f189fe572f0ed8cb1d064d772 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Tue, 25 Jul 2023 05:26:13 +0000 Subject: [PATCH 1269/1291] feat(algebraic_topology/dold_kan): tools for compatibilities, lemmas (#17923) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Joël Riou <37772949+joelriou@users.noreply.github.com> --- .../dold_kan/compatibility.lean | 125 +++++++++++++++++- 1 file changed, 119 insertions(+), 6 deletions(-) diff --git a/src/algebraic_topology/dold_kan/compatibility.lean b/src/algebraic_topology/dold_kan/compatibility.lean index b5afeedb24a51..3567efda0ccc9 100644 --- a/src/algebraic_topology/dold_kan/compatibility.lean +++ b/src/algebraic_topology/dold_kan/compatibility.lean @@ -77,7 +77,7 @@ lemma equivalence₁_inverse : (equivalence₁ hF).inverse = e'.inverse ⋙ eA.i def equivalence₁_counit_iso : (e'.inverse ⋙ eA.inverse) ⋙ F ≅ 𝟭 B' := calc (e'.inverse ⋙ eA.inverse) ⋙ F - ≅ (e'.inverse ⋙ eA.inverse) ⋙ (eA.functor ⋙ e'.functor) : iso_whisker_left _ hF.symm + ≅ (e'.inverse ⋙ eA.inverse) ⋙ (eA.functor ⋙ e'.functor) : iso_whisker_left _ hF.symm ... ≅ e'.inverse ⋙ (eA.inverse ⋙ eA.functor) ⋙ e'.functor : iso.refl _ ... ≅ e'.inverse ⋙ 𝟭 _ ⋙ e'.functor : iso_whisker_left _ (iso_whisker_right eA.counit_iso _) ... ≅ e'.inverse ⋙ e'.functor : iso.refl _ @@ -97,7 +97,7 @@ def equivalence₁_unit_iso : calc 𝟭 A ≅ eA.functor ⋙ eA.inverse : eA.unit_iso ... ≅ eA.functor ⋙ 𝟭 A' ⋙ eA.inverse : iso.refl _ ... ≅ eA.functor ⋙ (e'.functor ⋙ e'.inverse) ⋙ eA.inverse : - iso_whisker_left _ (iso_whisker_right e'.unit_iso _) + iso_whisker_left _ (iso_whisker_right e'.unit_iso _) ... ≅ (eA.functor ⋙ e'.functor) ⋙ (e'.inverse ⋙ eA.inverse) : iso.refl _ ... ≅ F ⋙ (e'.inverse ⋙ eA.inverse) : iso_whisker_right hF _ @@ -124,9 +124,9 @@ lemma equivalence₂_inverse : (equivalence₂ eB hF).inverse = def equivalence₂_counit_iso : (eB.functor ⋙ e'.inverse ⋙ eA.inverse) ⋙ (F ⋙ eB.inverse) ≅ 𝟭 B := calc (eB.functor ⋙ e'.inverse ⋙ eA.inverse) ⋙ (F ⋙ eB.inverse) - ≅ eB.functor ⋙ (e'.inverse ⋙ eA.inverse ⋙ F) ⋙ eB.inverse : iso.refl _ + ≅ eB.functor ⋙ (e'.inverse ⋙ eA.inverse ⋙ F) ⋙ eB.inverse : iso.refl _ ... ≅ eB.functor ⋙ 𝟭 _ ⋙ eB.inverse : - iso_whisker_left _ (iso_whisker_right (equivalence₁_counit_iso hF) _) + iso_whisker_left _ (iso_whisker_right (equivalence₁_counit_iso hF) _) ... ≅ eB.functor ⋙ eB.inverse : iso.refl _ ... ≅ 𝟭 B : eB.unit_iso.symm @@ -146,7 +146,7 @@ def equivalence₂_unit_iso : calc 𝟭 A ≅ F ⋙ e'.inverse ⋙ eA.inverse : equivalence₁_unit_iso hF ... ≅ F ⋙ 𝟭 B' ⋙ (e'.inverse ⋙ eA.inverse) : iso.refl _ ... ≅ F ⋙ (eB.inverse ⋙ eB.functor) ⋙ e'.inverse ⋙ eA.inverse : - iso_whisker_left _ (iso_whisker_right eB.counit_iso.symm _) + iso_whisker_left _ (iso_whisker_right eB.counit_iso.symm _) ... ≅ (F ⋙ eB.inverse) ⋙ (eB.functor ⋙ e'.inverse ⋙ eA.inverse) : iso.refl _ lemma equivalence₂_unit_iso_eq : @@ -169,7 +169,7 @@ begin letI : is_equivalence G := begin refine is_equivalence.of_iso _ (is_equivalence.of_equivalence (equivalence₂ eB hF).symm), calc eB.functor ⋙ e'.inverse ⋙ eA.inverse - ≅ (eB.functor ⋙ e'.inverse) ⋙ eA.inverse : iso.refl _ + ≅ (eB.functor ⋙ e'.inverse) ⋙ eA.inverse : iso.refl _ ... ≅ (G ⋙ eA.functor) ⋙ eA.inverse : iso_whisker_right hG _ ... ≅ G ⋙ 𝟭 A : iso_whisker_left _ eA.unit_iso.symm ... ≅ G : functor.right_unitor G, @@ -179,6 +179,119 @@ end lemma equivalence_functor : (equivalence hF hG).functor = F ⋙ eB.inverse := rfl +omit hG hF + +/-- The isomorphism `eB.functor ⋙ e'.inverse ⋙ e'.functor ≅ eB.functor` deduced +from the counit isomorphism of `e'`. -/ +@[simps hom_app] +def τ₀ : eB.functor ⋙ e'.inverse ⋙ e'.functor ≅ eB.functor := +calc eB.functor ⋙ e'.inverse ⋙ e'.functor + ≅ eB.functor ⋙ 𝟭 _ : iso_whisker_left _ e'.counit_iso +... ≅ eB.functor : functor.right_unitor _ + +include hF hG + +/-- The isomorphism `eB.functor ⋙ e'.inverse ⋙ e'.functor ≅ eB.functor` deduced +from the isomorphisms `hF : eA.functor ⋙ e'.functor ≅ F`, +`hG : eB.functor ⋙ e'.inverse ≅ G ⋙ eA.functor` and the datum of +an isomorphism `η : G ⋙ F ≅ eB.functor`. -/ +@[simps hom_app] +def τ₁ (η : G ⋙ F ≅ eB.functor) : + eB.functor ⋙ e'.inverse ⋙ e'.functor ≅ eB.functor := +calc eB.functor ⋙ e'.inverse ⋙ e'.functor + ≅ (eB.functor ⋙ e'.inverse) ⋙ e'.functor : iso.refl _ +... ≅ (G ⋙ eA.functor) ⋙ e'.functor : iso_whisker_right hG _ +... ≅ G ⋙ (eA.functor ⋙ e'.functor) : by refl +... ≅ G ⋙ F : iso_whisker_left _ hF +... ≅ eB.functor : η + +variables (η : G ⋙ F ≅ eB.functor) (hη : τ₀ = τ₁ hF hG η) + +omit hF hG +include η + +/-- The counit isomorphism of `equivalence`. -/ +@[simps] +def equivalence_counit_iso : G ⋙ (F ⋙ eB.inverse) ≅ 𝟭 B := +calc G ⋙ (F ⋙ eB.inverse) ≅ (G ⋙ F) ⋙ eB.inverse : iso.refl _ +... ≅ eB.functor ⋙ eB.inverse : iso_whisker_right η _ +... ≅ 𝟭 B : eB.unit_iso.symm + +variables {η hF hG} +include hη + +lemma equivalence_counit_iso_eq : + (equivalence hF hG).counit_iso = equivalence_counit_iso η := +begin + ext1, apply nat_trans.ext, ext Y, + dsimp [equivalence, equivalence_counit_iso, is_equivalence.of_equivalence], + simp only [equivalence₂_counit_iso_eq eB hF], + erw [nat_trans.id_app, nat_trans.id_app], + dsimp [equivalence₂, equivalence₁], + simp only [assoc, comp_id, F.map_id, id_comp, + equivalence₂_counit_iso_hom_app, ← eB.inverse.map_comp_assoc, + ← τ₀_hom_app, hη, τ₁_hom_app], + erw hF.inv.naturality_assoc, + congr' 2, + dsimp, + simp only [assoc, ← e'.functor.map_comp_assoc, eA.functor.map_comp, + equivalence.fun_inv_map, iso.inv_hom_id_app_assoc, hG.inv_hom_id_app], + dsimp, + rw [comp_id, eA.functor_unit_iso_comp, e'.functor.map_id, id_comp, hF.inv_hom_id_app_assoc], +end + +omit hη η eB +include hF + +variable (hF) + +/-- The isomorphism `eA.functor ≅ F ⋙ e'.inverse` deduced from the +unit isomorphism of `e'` and the isomorphism `hF : eA.functor ⋙ e'.functor ≅ F`. -/ +@[simps] +def υ : eA.functor ≅ F ⋙ e'.inverse := +calc eA.functor ≅ eA.functor ⋙ 𝟭 A' : (functor.left_unitor _).symm +... ≅ eA.functor ⋙ (e'.functor ⋙ e'.inverse) : iso_whisker_left _ e'.unit_iso +... ≅ (eA.functor ⋙ e'.functor) ⋙ e'.inverse : iso.refl _ +... ≅ F ⋙ e'.inverse : iso_whisker_right hF _ + +variables (ε : eA.functor ≅ F ⋙ e'.inverse) (hε : υ hF = ε) + +include ε hG +omit hF + +variable (hG) + +/-- The unit isomorphism of `equivalence`. -/ +@[simps] +def equivalence_unit_iso : 𝟭 A ≅ (F ⋙ eB.inverse) ⋙ G := +calc 𝟭 A ≅ eA.functor ⋙ eA.inverse : eA.unit_iso +... ≅ (F ⋙ e'.inverse) ⋙ eA.inverse : iso_whisker_right ε _ +... ≅ F ⋙ 𝟭 B' ⋙ e'.inverse ⋙ eA.inverse : iso.refl _ +... ≅ F ⋙ (eB.inverse ⋙ eB.functor) ⋙ (e'.inverse ⋙ eA.inverse) : + iso_whisker_left _ (iso_whisker_right eB.counit_iso.symm _) +... ≅ (F ⋙ eB.inverse) ⋙ (eB.functor ⋙ e'.inverse) ⋙ eA.inverse : iso.refl _ +... ≅ (F ⋙ eB.inverse) ⋙ (G ⋙ eA.functor) ⋙ eA.inverse : + iso_whisker_left _ (iso_whisker_right hG _) +... ≅ (F ⋙ eB.inverse ⋙ G) ⋙ (eA.functor ⋙ eA.inverse) : iso.refl _ +... ≅ (F ⋙ eB.inverse ⋙ G) ⋙ 𝟭 A : iso_whisker_left _ eA.unit_iso.symm +... ≅ (F ⋙ eB.inverse) ⋙ G : iso.refl _ + +include hε +variables {ε hF hG} + +lemma equivalence_unit_iso_eq : + (equivalence hF hG).unit_iso = equivalence_unit_iso hG ε := +begin + ext1, apply nat_trans.ext, ext X, + dsimp [equivalence, iso.refl, nat_iso.hcomp, is_equivalence.inverse, + is_equivalence.of_equivalence], + erw [nat_trans.id_app, id_comp, G.map_id, comp_id, comp_id], + simp only [equivalence₂_unit_iso_eq eB hF, equivalence₂_unit_iso_hom_app], + dsimp [equivalence₂, equivalence₁], + simp only [assoc, equivalence_unit_iso_hom_app, nat_iso.cancel_nat_iso_hom_left, + ← eA.inverse.map_comp_assoc, ← hε, υ_hom_app], +end + end compatibility end dold_kan From d11f435d4e34a6cea0a1797d6b625b0c170be845 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 26 Jul 2023 13:39:42 +0000 Subject: [PATCH 1270/1291] feat(linear_algebra/quadratic_form): some work from 'Clifford Algebras and Spinor Norms Over a Commutative Ring' (#18447) --- docs/references.bib | 12 ++ src/linear_algebra/quadratic_form/basic.lean | 2 +- src/linear_algebra/quadratic_form/dual.lean | 151 +++++++++++++++++++ 3 files changed, 164 insertions(+), 1 deletion(-) create mode 100644 src/linear_algebra/quadratic_form/dual.lean diff --git a/docs/references.bib b/docs/references.bib index 03a341208d7bb..ece945a541bf9 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -1317,6 +1317,18 @@ @Article{ huneke2002 url = {https://doi.org/10.1080/00029890.2002.11919853} } +@InProceedings{ hyman1973, + author = "Bass, Hyman", + editor = "Bass, Hyman", + title = "Unitary algebraic K-theory", + booktitle = "Hermitian K-Theory and Geometric Applications", + year = "1973", + publisher = "Springer Berlin Heidelberg", + address = "Berlin, Heidelberg", + pages = "57--265", + isbn = "978-3-540-37773-3" +} + @Book{ iordanescu2003, author = {Radu {Iord\u{a}nescu}}, title = {{Jordan structures in geometry and physics. With an diff --git a/src/linear_algebra/quadratic_form/basic.lean b/src/linear_algebra/quadratic_form/basic.lean index cea7b02947391..663a7df5eafb5 100644 --- a/src/linear_algebra/quadratic_form/basic.lean +++ b/src/linear_algebra/quadratic_form/basic.lean @@ -706,7 +706,7 @@ variables [invertible (2 : R₁)] /-- `associated` is the linear map that sends a quadratic form over a commutative ring to its associated symmetric bilinear form. -/ -abbreviation associated : quadratic_form R₁ M →ₗ[R₁] bilin_form R₁ M := +@[reducible] def associated : quadratic_form R₁ M →ₗ[R₁] bilin_form R₁ M := associated_hom R₁ @[simp] lemma associated_lin_mul_lin (f g : M →ₗ[R₁] R₁) : diff --git a/src/linear_algebra/quadratic_form/dual.lean b/src/linear_algebra/quadratic_form/dual.lean new file mode 100644 index 0000000000000..68ed89fab74cb --- /dev/null +++ b/src/linear_algebra/quadratic_form/dual.lean @@ -0,0 +1,151 @@ +/- +Copyright (c) 2023 Eric Wieser. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Eric Wieser +-/ + +import linear_algebra.quadratic_form.isometry +import linear_algebra.quadratic_form.prod +/-! +# Quadratic form structures related to `module.dual` + +## Main definitions + +* `bilin_form.dual_prod R M`, the bilinear form on `(f, x) : module.dual R M × M` defined as + `f x`. +* `quadratic_form.dual_prod R M`, the quadratic form on `(f, x) : module.dual R M × M` defined as + `f x`. +* `quadratic_form.to_dual_prod : M × M →ₗ[R] module.dual R M × M` a form-preserving map from + `(Q.prod $ -Q)` to `quadratic_form.dual_prod R M`. Note that we do not have the morphism + version of `quadratic_form.isometry`, so for now this is stated without full bundling. + +-/ + +variables (R M N : Type*) + +namespace bilin_form + +section semiring +variables [comm_semiring R] [add_comm_monoid M] [module R M] + +/-- The symmetric bilinear form on `module.dual R M × M` defined as +`B (f, x) (g, y) = f y + g x`. -/ +@[simps] +def dual_prod : bilin_form R (module.dual R M × M) := +linear_map.to_bilin $ + (linear_map.applyₗ.comp (linear_map.snd R (module.dual R M) M)).compl₂ + (linear_map.fst R (module.dual R M) M) + + ((linear_map.applyₗ.comp (linear_map.snd R (module.dual R M) M)).compl₂ + (linear_map.fst R (module.dual R M) M)).flip + +lemma is_symm_dual_prod : (dual_prod R M).is_symm := +λ x y, add_comm _ _ + +end semiring + +section ring +variables [comm_ring R] [add_comm_group M] [module R M] + +lemma nondenerate_dual_prod : + (dual_prod R M).nondegenerate ↔ function.injective (module.dual.eval R M) := +begin + classical, + rw nondegenerate_iff_ker_eq_bot, + rw linear_map.ker_eq_bot, + let e := linear_equiv.prod_comm R _ _ + ≪≫ₗ module.dual_prod_dual_equiv_dual R (module.dual R M) M, + let h_d := e.symm.to_linear_map.comp (dual_prod R M).to_lin, + refine (function.injective.of_comp_iff e.symm.injective (dual_prod R M).to_lin).symm.trans _, + rw [←linear_equiv.coe_to_linear_map, ←linear_map.coe_comp], + change function.injective h_d ↔ _, + have : h_d = linear_map.prod_map (linear_map.id) (module.dual.eval R M), + { refine linear_map.ext (λ x, prod.ext _ _), + { ext, + dsimp [h_d, module.dual.eval, linear_equiv.prod_comm], + simp }, + { ext, + dsimp [h_d, module.dual.eval, linear_equiv.prod_comm], + simp } }, + rw [this, linear_map.coe_prod_map], + refine prod.map_injective.trans _, + exact and_iff_right function.injective_id, +end + +end ring + +end bilin_form + +namespace quadratic_form + +section semiring +variables [comm_semiring R] [add_comm_monoid M] [add_comm_monoid N] [module R M] [module R N] + +/-- The quadratic form on `module.dual R M × M` defined as `Q (f, x) = f x`. -/ +@[simps] +def dual_prod : quadratic_form R (module.dual R M × M) := +{ to_fun := λ p, p.1 p.2, + to_fun_smul := λ a p, by rw [prod.smul_fst, prod.smul_snd, linear_map.smul_apply, + linear_map.map_smul, smul_eq_mul, smul_eq_mul, mul_assoc], + exists_companion' := ⟨bilin_form.dual_prod R M, λ p q, begin + rw [bilin_form.dual_prod_apply, prod.fst_add, prod.snd_add, linear_map.add_apply, map_add, + map_add, add_right_comm _ (q.1 q.2), add_comm (q.1 p.2) (p.1 q.2), ←add_assoc, ←add_assoc], + end⟩ } + +@[simp] +lemma _root_.bilin_form.dual_prod.to_quadratic_form : + (bilin_form.dual_prod R M).to_quadratic_form = 2 • dual_prod R M := +ext $ λ a, (two_nsmul _).symm + +variables {R M N} + +/-- Any module isomorphism induces a quadratic isomorphism between the corresponding `dual_prod.` -/ +@[simps] +def dual_prod_isometry (f : M ≃ₗ[R] N) : + (dual_prod R M).isometry (dual_prod R N) := +{ to_linear_equiv := f.dual_map.symm.prod f, + map_app' := λ x, fun_like.congr_arg x.fst $ f.symm_apply_apply _ } + +/-- `quadratic_form.dual_prod` commutes (isometrically) with `quadratic_form.prod`. -/ +@[simps] +def dual_prod_prod_isometry : + (dual_prod R (M × N)).isometry ((dual_prod R M).prod (dual_prod R N)) := +{ to_linear_equiv := + ((module.dual_prod_dual_equiv_dual R M N).symm.prod (linear_equiv.refl R (M × N))) + ≪≫ₗ linear_equiv.prod_prod_prod_comm R _ _ M N, + map_app' := λ m, + (m.fst.map_add _ _).symm.trans $ fun_like.congr_arg m.fst $ prod.ext (add_zero _) (zero_add _) } + +end semiring + +section ring +variables [comm_ring R] [add_comm_group M] [module R M] + +variables {R M} + +/-- The isometry sending `(Q.prod $ -Q)` to `(quadratic_form.dual_prod R M)`. + +This is `σ` from Proposition 4.8, page 84 of +[*Hermitian K-Theory and Geometric Applications*][hyman1973]; though we swap the order of the pairs. +-/ +@[simps] +def to_dual_prod (Q : quadratic_form R M) [invertible (2 : R)] : + M × M →ₗ[R] module.dual R M × M := +linear_map.prod + (Q.associated.to_lin.comp (linear_map.fst _ _ _) + + Q.associated.to_lin.comp (linear_map.snd _ _ _)) + ((linear_map.fst _ _ _ - linear_map.snd _ _ _)) + +lemma to_dual_prod_isometry [invertible (2 : R)] (Q : quadratic_form R M) (x : M × M) : + quadratic_form.dual_prod R M (to_dual_prod Q x) = (Q.prod $ -Q) x := +begin + dsimp only [to_dual_prod, associated, associated_hom], + dsimp, + simp [polar_comm _ x.1 x.2, ←sub_add, mul_sub, sub_mul, smul_sub, submonoid.smul_def, + ←sub_eq_add_neg (Q x.1) (Q x.2)], +end + +-- TODO: show that `to_dual_prod` is an equivalence + +end ring + +end quadratic_form From fe18deda804e30c594e75a6e5fe0f7d14695289f Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 27 Jul 2023 10:40:44 +0000 Subject: [PATCH 1271/1291] feat(analysis/normed_space/continuous_linear_map): generalize typeclass assumptions (#19108) --- .../normed_space/continuous_linear_map.lean | 81 ++++++++++--------- 1 file changed, 41 insertions(+), 40 deletions(-) diff --git a/src/analysis/normed_space/continuous_linear_map.lean b/src/analysis/normed_space/continuous_linear_map.lean index a46af72230230..1ca9737ceac0a 100644 --- a/src/analysis/normed_space/continuous_linear_map.lean +++ b/src/analysis/normed_space/continuous_linear_map.lean @@ -35,14 +35,14 @@ open_locale nnreal variables {𝕜 𝕜₂ E F G : Type*} -variables [normed_field 𝕜] [normed_field 𝕜₂] /-! # General constructions -/ -section seminormed +section seminormed_add_comm_group +variables [ring 𝕜] [ring 𝕜₂] variables [seminormed_add_comm_group E] [seminormed_add_comm_group F] [seminormed_add_comm_group G] -variables [normed_space 𝕜 E] [normed_space 𝕜₂ F] [normed_space 𝕜 G] +variables [module 𝕜 E] [module 𝕜₂ F] [module 𝕜 G] variables {σ : 𝕜 →+* 𝕜₂} (f : E →ₛₗ[σ] F) /-- Construct a continuous linear map from a linear map and a bound on this linear map. @@ -51,13 +51,6 @@ The fact that the norm of the continuous linear map is then controlled is given def linear_map.mk_continuous (C : ℝ) (h : ∀x, ‖f x‖ ≤ C * ‖x‖) : E →SL[σ] F := ⟨f, add_monoid_hom_class.continuous_of_bound f C h⟩ -/-- Reinterpret a linear map `𝕜 →ₗ[𝕜] E` as a continuous linear map. This construction -is generalized to the case of any finite dimensional domain -in `linear_map.to_continuous_linear_map`. -/ -def linear_map.to_continuous_linear_map₁ (f : 𝕜 →ₗ[𝕜] E) : 𝕜 →L[𝕜] E := -f.mk_continuous (‖f 1‖) $ λ x, le_of_eq $ -by { conv_lhs { rw ← mul_one x }, rw [← smul_eq_mul, f.map_smul, norm_smul, mul_comm] } - /-- Construct a continuous linear map from a linear map and the existence of a bound on this linear map. If you have an explicit bound, use `linear_map.mk_continuous` instead, as a norm estimate will follow automatically in `linear_map.mk_continuous_norm_le`. -/ @@ -89,14 +82,6 @@ add_monoid_hom_class.continuous_of_bound φ C h_bound @[simp] lemma linear_map.mk_continuous_of_exists_bound_apply (h : ∃C, ∀x, ‖f x‖ ≤ C * ‖x‖) (x : E) : f.mk_continuous_of_exists_bound h x = f x := rfl -@[simp] lemma linear_map.to_continuous_linear_map₁_coe (f : 𝕜 →ₗ[𝕜] E) : - (f.to_continuous_linear_map₁ : 𝕜 →ₗ[𝕜] E) = f := -rfl - -@[simp] lemma linear_map.to_continuous_linear_map₁_apply (f : 𝕜 →ₗ[𝕜] E) (x) : - f.to_continuous_linear_map₁ x = f x := -rfl - namespace continuous_linear_map theorem antilipschitz_of_bound (f : E →SL[σ] F) {K : ℝ≥0} (h : ∀ x, ‖x‖ ≤ K * ‖f x‖) : @@ -125,11 +110,34 @@ def linear_equiv.to_continuous_linear_equiv_of_bounds (e : E ≃ₛₗ[σ] F) (C end -end seminormed +end seminormed_add_comm_group + +section seminormed_bounded + +variables [semi_normed_ring 𝕜] [ring 𝕜₂] [seminormed_add_comm_group E] +variables [module 𝕜 E] [has_bounded_smul 𝕜 E] + +/-- Reinterpret a linear map `𝕜 →ₗ[𝕜] E` as a continuous linear map. This construction +is generalized to the case of any finite dimensional domain +in `linear_map.to_continuous_linear_map`. -/ +def linear_map.to_continuous_linear_map₁ (f : 𝕜 →ₗ[𝕜] E) : 𝕜 →L[𝕜] E := +f.mk_continuous (‖f 1‖) $ λ x, +by { conv_lhs { rw ← mul_one x }, rw [← smul_eq_mul, f.map_smul, mul_comm],exact norm_smul_le _ _ } + +@[simp] lemma linear_map.to_continuous_linear_map₁_coe (f : 𝕜 →ₗ[𝕜] E) : + (f.to_continuous_linear_map₁ : 𝕜 →ₗ[𝕜] E) = f := +rfl + +@[simp] lemma linear_map.to_continuous_linear_map₁_apply (f : 𝕜 →ₗ[𝕜] E) (x) : + f.to_continuous_linear_map₁ x = f x := +rfl + +end seminormed_bounded section normed -variables [normed_add_comm_group E] [normed_add_comm_group F] [normed_space 𝕜 E] [normed_space 𝕜₂ F] +variables [ring 𝕜] [ring 𝕜₂] +variables [normed_add_comm_group E] [normed_add_comm_group F] [module 𝕜 E] [module 𝕜₂ F] variables {σ : 𝕜 →+* 𝕜₂} (f g : E →SL[σ] F) (x y z : E) theorem continuous_linear_map.uniform_embedding_of_bound {K : ℝ≥0} (hf : ∀ x, ‖x‖ ≤ K * ‖f x‖) : @@ -142,8 +150,9 @@ end normed section seminormed +variables [ring 𝕜] [ring 𝕜₂] variables [seminormed_add_comm_group E] [seminormed_add_comm_group F] -variables [normed_space 𝕜 E] [normed_space 𝕜₂ F] +variables [module 𝕜 E] [module 𝕜₂ F] variables {σ : 𝕜 →+* 𝕜₂} (f : E →ₛₗ[σ] F) /-- A (semi-)linear map which is a homothety is a continuous linear map. @@ -180,42 +189,34 @@ end seminormed /-! ## The span of a single vector -/ -section seminormed - -variables [seminormed_add_comm_group E] [normed_space 𝕜 E] - namespace continuous_linear_map +variables (𝕜) -variable (𝕜) +section seminormed +variables [normed_division_ring 𝕜] [seminormed_add_comm_group E] [module 𝕜 E] [has_bounded_smul 𝕜 E] lemma to_span_singleton_homothety (x : E) (c : 𝕜) : ‖linear_map.to_span_singleton 𝕜 E x c‖ = ‖x‖ * ‖c‖ := by {rw mul_comm, exact norm_smul _ _} -end continuous_linear_map +end seminormed -section +end continuous_linear_map namespace continuous_linear_equiv - variable (𝕜) +section seminormed +variables [normed_division_ring 𝕜] [seminormed_add_comm_group E] [module 𝕜 E] [has_bounded_smul 𝕜 E] + lemma to_span_nonzero_singleton_homothety (x : E) (h : x ≠ 0) (c : 𝕜) : ‖linear_equiv.to_span_nonzero_singleton 𝕜 E x h c‖ = ‖x‖ * ‖c‖ := continuous_linear_map.to_span_singleton_homothety _ _ _ -end continuous_linear_equiv - -end - end seminormed section normed - -variables [normed_add_comm_group E] [normed_space 𝕜 E] - -namespace continuous_linear_equiv -variable (𝕜) +variables [normed_field 𝕜] [normed_add_comm_group E] [normed_space 𝕜 E] /-- Given a nonzero element `x` of a normed space `E₁` over a field `𝕜`, the natural continuous linear equivalence from `E₁` to the span of `x`.-/ @@ -246,6 +247,6 @@ noncomputable def coord (x : E) (h : x ≠ 0) : (𝕜 ∙ x) →L[𝕜] 𝕜 := (coord 𝕜 x h) (⟨x, submodule.mem_span_singleton_self x⟩ : 𝕜 ∙ x) = 1 := linear_equiv.coord_self 𝕜 E x h -end continuous_linear_equiv - end normed + +end continuous_linear_equiv From e90e0a6119d01e3ef1703b3038b555785a998285 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 28 Jul 2023 07:23:47 +0000 Subject: [PATCH 1272/1291] feat(tactic/positivity): Extension for `ite` (#17650) Add `positivity_ite`, an extension for `ite`. Co-authored-by: Jireh Loreaux --- src/tactic/positivity.lean | 46 ++++++++++++++++++++++++++++++++++++++ test/positivity.lean | 12 ++++++++++ 2 files changed, 58 insertions(+) diff --git a/src/tactic/positivity.lean b/src/tactic/positivity.lean index 6544387973144..bec4bfd4a8cb8 100644 --- a/src/tactic/positivity.lean +++ b/src/tactic/positivity.lean @@ -352,6 +352,52 @@ variables {ι α R : Type*} /-! ### `positivity` extensions for particular arithmetic operations -/ +section ite +variables [has_zero α] {p : Prop} [decidable p] {a b : α} + +private lemma ite_pos [has_lt α] (ha : 0 < a) (hb : 0 < b) : 0 < ite p a b := +by by_cases p; simp [*] + +private lemma ite_nonneg [has_le α] (ha : 0 ≤ a) (hb : 0 ≤ b) : 0 ≤ ite p a b := +by by_cases p; simp [*] + +private lemma ite_nonneg_of_pos_of_nonneg [preorder α] (ha : 0 < a) (hb : 0 ≤ b) : 0 ≤ ite p a b := +ite_nonneg ha.le hb + +private lemma ite_nonneg_of_nonneg_of_pos [preorder α] (ha : 0 ≤ a) (hb : 0 < b) : 0 ≤ ite p a b := +ite_nonneg ha hb.le + +private lemma ite_ne_zero (ha : a ≠ 0) (hb : b ≠ 0) : ite p a b ≠ 0 := by by_cases p; simp [*] + +private lemma ite_ne_zero_of_pos_of_ne_zero [preorder α] (ha : 0 < a) (hb : b ≠ 0) : + ite p a b ≠ 0 := +ite_ne_zero ha.ne' hb + +private lemma ite_ne_zero_of_ne_zero_of_pos [preorder α] (ha : a ≠ 0) (hb : 0 < b) : + ite p a b ≠ 0 := +ite_ne_zero ha hb.ne' + +end ite + +/-- Extension for the `positivity` tactic: the `if then else` of two numbers is +positive/nonnegative/nonzero if both are. -/ +@[positivity] +meta def positivity_ite : expr → tactic strictness +| e@`(@ite %%typ %%p %%hp %%a %%b) := do + strictness_a ← core a, + strictness_b ← core b, + match strictness_a, strictness_b with + | positive pa, positive pb := positive <$> mk_app ``ite_pos [pa, pb] + | positive pa, nonnegative pb := nonnegative <$> mk_app ``ite_nonneg_of_pos_of_nonneg [pa, pb] + | nonnegative pa, positive pb := nonnegative <$> mk_app ``ite_nonneg_of_nonneg_of_pos [pa, pb] + | nonnegative pa, nonnegative pb := nonnegative <$> mk_app ``ite_nonneg [pa, pb] + | positive pa, nonzero pb := nonzero <$> to_expr ``(ite_ne_zero_of_pos_of_ne_zero %%pa %%pb) + | nonzero pa, positive pb := nonzero <$> to_expr ``(ite_ne_zero_of_ne_zero_of_pos %%pa %%pb) + | nonzero pa, nonzero pb := nonzero <$> to_expr ``(ite_ne_zero %%pa %%pb) + | sa@_, sb@ _ := positivity_fail e a b sa sb + end +| e := pp e >>= fail ∘ format.bracket "The expression `" "` isn't of the form `ite p a b`" + section linear_order variables [linear_order R] {a b c : R} diff --git a/test/positivity.lean b/test/positivity.lean index f4dc7425036ef..007d449a37bf3 100644 --- a/test/positivity.lean +++ b/test/positivity.lean @@ -80,6 +80,18 @@ example [has_zero α] [preorder α] {a : α} (ha : 0 < a) : 0 ≤ const ι a := example [has_zero α] [preorder α] {a : α} (ha : 0 ≤ a) : 0 ≤ const ι a := by positivity example [nonempty ι] [has_zero α] [preorder α] {a : α} (ha : 0 < a) : 0 < const ι a := by positivity +section ite +variables {p : Prop} [decidable p] {a b : ℤ} + +example (ha : 0 < a) (hb : 0 < b) : 0 < ite p a b := by positivity +example (ha : 0 < a) (hb : 0 ≤ b) : 0 ≤ ite p a b := by positivity +example (ha : 0 ≤ a) (hb : 0 < b) : 0 ≤ ite p a b := by positivity +example (ha : 0 < a) (hb : b ≠ 0) : ite p a b ≠ 0 := by positivity +example (ha : a ≠ 0) (hb : 0 < b) : ite p a b ≠ 0 := by positivity +example (ha : a ≠ 0) (hb : b ≠ 0) : ite p a b ≠ 0 := by positivity + +end ite + example {a b : ℚ} (ha : 0 < a) (hb : 0 < b) : 0 < min a b := by positivity example {a b : ℚ} (ha : 0 < a) (hb : 0 ≤ b) : 0 ≤ min a b := by positivity example {a b : ℚ} (ha : 0 ≤ a) (hb : 0 < b) : 0 ≤ min a b := by positivity From 8900d545017cd21961daa2a1734bb658ef52c618 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Fri, 28 Jul 2023 10:11:06 +0000 Subject: [PATCH 1273/1291] feat(set_theory/game/pgame): small sets of pre-games / games / surreals are bounded (#15260) --- src/set_theory/game/basic.lean | 10 ++++++ src/set_theory/game/pgame.lean | 53 +++++++++++++++++++++++++++++++ src/set_theory/surreal/basic.lean | 28 ++++++++++++++++ 3 files changed, 91 insertions(+) diff --git a/src/set_theory/game/basic.lean b/src/set_theory/game/basic.lean index e98541fca9c39..d1f574448fbed 100644 --- a/src/set_theory/game/basic.lean +++ b/src/set_theory/game/basic.lean @@ -128,6 +128,16 @@ instance ordered_add_comm_group : ordered_add_comm_group game := ..game.add_comm_group_with_one, ..game.partial_order } +/-- A small set `s` of games is bounded above. -/ +lemma bdd_above_of_small (s : set game.{u}) [small.{u} s] : bdd_above s := +⟨_, λ i hi, by simpa using pgame.le_iff_game_le.1 + (upper_bound_mem_upper_bounds _ (set.mem_image_of_mem quotient.out hi))⟩ + +/-- A small set `s` of games is bounded below. -/ +lemma bdd_below_of_small (s : set game.{u}) [small.{u} s] : bdd_below s := +⟨_, λ i hi, by simpa using pgame.le_iff_game_le.1 + (lower_bound_mem_lower_bounds _ (set.mem_image_of_mem quotient.out hi))⟩ + end game namespace pgame diff --git a/src/set_theory/game/pgame.lean b/src/set_theory/game/pgame.lean index 1cfa0e55c037c..b6a9580b9cbb2 100644 --- a/src/set_theory/game/pgame.lean +++ b/src/set_theory/game/pgame.lean @@ -6,6 +6,7 @@ Authors: Reid Barton, Mario Carneiro, Isabel Longbottom, Scott Morrison import data.fin.basic import data.list.basic import logic.relation +import logic.small.basic import order.game_add /-! @@ -491,6 +492,58 @@ lemma left_response_spec {x : pgame} (h : 0 ≤ x) (j : x.right_moves) : 0 ≤ (x.move_right j).move_left (left_response h j) := classical.some_spec $ (zero_le.1 h) j +/-- An explicit upper bound for a family of pre-games, whose left moves are the union of the left +moves of all the pre-games in the family. -/ +def upper_bound {ι : Type u} (f : ι → pgame.{u}) : pgame := +⟨Σ i, (f i).left_moves, pempty, λ x, move_left _ x.2, pempty.elim⟩ + +instance upper_bound_right_moves_empty {ι : Type u} (f : ι → pgame.{u}) : + is_empty (upper_bound f).right_moves := +pempty.is_empty + +theorem le_upper_bound {ι : Type u} (f : ι → pgame.{u}) (i : ι) : f i ≤ upper_bound f := +begin + rw [upper_bound, le_iff_forall_lf], + dsimp, + simp only [and_true, is_empty.forall_iff], + exact λ j, @move_left_lf (upper_bound f) ⟨i, j⟩ +end + +lemma upper_bound_mem_upper_bounds (s : set pgame.{u}) [small.{u} s] : + upper_bound (subtype.val ∘ (equiv_shrink s).symm) ∈ upper_bounds s := +λ i hi, by simpa using + le_upper_bound (subtype.val ∘ (equiv_shrink s).symm) (equiv_shrink s ⟨i, hi⟩) + +/-- A small set `s` of pre-games is bounded above. -/ +lemma bdd_above_of_small (s : set pgame.{u}) [small.{u} s] : bdd_above s := +⟨_, upper_bound_mem_upper_bounds s⟩ + +/-- An explicit lower bound for a family of pre-games, whose right moves are the union of the right +moves of all the pre-games in the family. -/ +def lower_bound {ι : Type u} (f : ι → pgame.{u}) : pgame := +⟨pempty, Σ i, (f i).right_moves, pempty.elim, λ x, move_right _ x.2⟩ + +instance lower_bound_left_moves_empty {ι : Type u} (f : ι → pgame.{u}) : + is_empty (lower_bound f).left_moves := +pempty.is_empty + +theorem lower_bound_le {ι : Type u} (f : ι → pgame.{u}) (i : ι) : lower_bound f ≤ f i := +begin + rw [lower_bound, le_iff_forall_lf], + dsimp, + simp only [is_empty.forall_iff, true_and], + exact λ j, @lf_move_right (lower_bound f) ⟨i, j⟩ +end + +lemma lower_bound_mem_lower_bounds (s : set pgame.{u}) [small.{u} s] : + lower_bound (subtype.val ∘ (equiv_shrink s).symm) ∈ lower_bounds s := +λ i hi, by simpa using + lower_bound_le (subtype.val ∘ (equiv_shrink s).symm) (equiv_shrink s ⟨i, hi⟩) + +/-- A small set `s` of pre-games is bounded below. -/ +lemma bdd_below_of_small (s : set pgame.{u}) [small.{u} s] : bdd_below s := +⟨_, lower_bound_mem_lower_bounds s⟩ + /-- The equivalence relation on pre-games. Two pre-games `x`, `y` are equivalent if `x ≤ y` and `y ≤ x`. diff --git a/src/set_theory/surreal/basic.lean b/src/set_theory/surreal/basic.lean index 6a9d8bf27327a..fe2df6135ac01 100644 --- a/src/set_theory/surreal/basic.lean +++ b/src/set_theory/surreal/basic.lean @@ -328,6 +328,34 @@ theorem zero_to_game : to_game 0 = 0 := rfl @[simp] theorem one_to_game : to_game 1 = 1 := rfl @[simp] theorem nat_to_game : ∀ n : ℕ, to_game n = n := map_nat_cast' _ one_to_game +theorem upper_bound_numeric {ι : Type u} {f : ι → pgame.{u}} (H : ∀ i, (f i).numeric) : + (upper_bound f).numeric := +numeric_of_is_empty_right_moves _ $ λ i, (H _).move_left _ + +theorem lower_bound_numeric {ι : Type u} {f : ι → pgame.{u}} (H : ∀ i, (f i).numeric) : + (lower_bound f).numeric := +numeric_of_is_empty_left_moves _ $ λ i, (H _).move_right _ + +/-- A small set `s` of surreals is bounded above. -/ +lemma bdd_above_of_small (s : set surreal.{u}) [small.{u} s] : bdd_above s := +begin + let g := subtype.val ∘ quotient.out ∘ subtype.val ∘ (equiv_shrink s).symm, + refine ⟨mk (upper_bound g) (upper_bound_numeric $ λ i, subtype.prop _), λ i hi, _⟩, + rw ←quotient.out_eq i, + show i.out.1 ≤ _, + simpa [g] using le_upper_bound g (equiv_shrink s ⟨i, hi⟩) +end + +/-- A small set `s` of surreals is bounded below. -/ +lemma bdd_below_of_small (s : set surreal.{u}) [small.{u} s] : bdd_below s := +begin + let g := subtype.val ∘ quotient.out ∘ subtype.val ∘ (equiv_shrink s).symm, + refine ⟨mk (lower_bound g) (lower_bound_numeric $ λ i, subtype.prop _), λ i hi, _⟩, + rw ←quotient.out_eq i, + show _ ≤ i.out.1, + simpa [g] using lower_bound_le g (equiv_shrink s ⟨i, hi⟩) +end + end surreal open surreal From 188a411e916e1119e502dbe35b8b475716362401 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9mi=20Bottinelli?= Date: Fri, 28 Jul 2023 10:11:08 +0000 Subject: [PATCH 1274/1291] feat(combinatorics/quiver/covering): Definition of coverings and unique lifting of paths (#17828) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Antoine Labelle Co-authored-by: Rémi Bottinelli Co-authored-by: Kyle Miller Co-authored-by: antoinelab01 <66086247+antoinelab01@users.noreply.github.com> Co-authored-by: Antoine Labelle --- src/combinatorics/quiver/covering.lean | 266 +++++++++++++++++++++++++ 1 file changed, 266 insertions(+) create mode 100644 src/combinatorics/quiver/covering.lean diff --git a/src/combinatorics/quiver/covering.lean b/src/combinatorics/quiver/covering.lean new file mode 100644 index 0000000000000..18c98970d2cfd --- /dev/null +++ b/src/combinatorics/quiver/covering.lean @@ -0,0 +1,266 @@ +/- +Copyright (c) 2022 Antoine Labelle, Rémi Bottinelli. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Antoine Labelle, Rémi Bottinelli +-/ +import combinatorics.quiver.cast +import combinatorics.quiver.symmetric +import data.sigma.basic +import logic.equiv.basic + +/-! +# Covering + +This file defines coverings of quivers as prefunctors that are bijective on the +so-called stars and costars at each vertex of the domain. + +## Main definitions + +* `quiver.star u` is the type of all arrows with source `u`; +* `quiver.costar u` is the type of all arrows with target `u`; +* `prefunctor.star φ u` is the obvious function `star u → star (φ.obj u)`; +* `prefunctor.costar φ u` is the obvious function `costar u → costar (φ.obj u)`; +* `prefunctor.is_covering φ` means that `φ.star u` and `φ.costar u` are bijections for all `u`; +* `quiver.star_path u` is the type of all paths with source `u`; +* `prefunctor.star_path u` is the obvious function `star_path u → star_path (φ.obj u)`. + +## Main statements + +* `prefunctor.is_covering.path_star_bijective` states that if `φ` is a covering, + then `φ.star_path u` is a bijection for all `u`. + In other words, every path in the codomain of `φ` lifts uniquely to its domain. + +## TODO + +Clean up the namespaces by renaming `prefunctor` to `quiver.prefunctor`. + +## Tags + +Cover, covering, quiver, path, lift +-/ + +open function quiver + +universes u v w + +variables {U : Type*} [quiver.{u+1} U] + {V : Type*} [quiver.{v+1} V] (φ : U ⥤q V) + {W : Type*} [quiver.{w+1} W] (ψ : V ⥤q W) + +/-- The `quiver.star` at a vertex is the collection of arrows whose source is the vertex. +The type `quiver.star u` is defined to be `Σ (v : U), (u ⟶ v)`. -/ +@[reducible] def quiver.star (u : U) := Σ (v : U), (u ⟶ v) + +/-- Constructor for `quiver.star`. Defined to be `sigma.mk`. -/ +@[reducible] protected def quiver.star.mk {u v : U} (f : u ⟶ v) : quiver.star u := ⟨_, f⟩ + +/-- The `quiver.costar` at a vertex is the collection of arrows whose target is the vertex. +The type `quiver.costar v` is defined to be `Σ (u : U), (u ⟶ v)`. -/ +@[reducible] def quiver.costar (v : U) := Σ (u : U), (u ⟶ v) + +/-- Constructor for `quiver.costar`. Defined to be `sigma.mk`. -/ +@[reducible] protected def quiver.costar.mk {u v : U} (f : u ⟶ v) : quiver.costar v := ⟨_, f⟩ + +/-- A prefunctor induces a map of `quiver.star` at every vertex. -/ +@[simps] def prefunctor.star (u : U) : quiver.star u → quiver.star (φ.obj u) := +λ F, quiver.star.mk (φ.map F.2) + +/-- A prefunctor induces a map of `quiver.costar` at every vertex. -/ +@[simps] def prefunctor.costar (u : U) : quiver.costar u → quiver.costar (φ.obj u) := +λ F, quiver.costar.mk (φ.map F.2) + +@[simp] lemma prefunctor.star_apply {u v : U} (e : u ⟶ v) : + φ.star u (quiver.star.mk e) = quiver.star.mk (φ.map e) := rfl + +@[simp] lemma prefunctor.costar_apply {u v : U} (e : u ⟶ v) : + φ.costar v (quiver.costar.mk e) = quiver.costar.mk (φ.map e) := rfl + +lemma prefunctor.star_comp (u : U) : + (φ ⋙q ψ).star u = (ψ.star (φ.obj u)) ∘ ((φ.star) u) := rfl + +lemma prefunctor.costar_comp (u : U) : + (φ ⋙q ψ).costar u = (ψ.costar (φ.obj u)) ∘ ((φ.costar) u) := rfl + +/-- A prefunctor is a covering of quivers if it defines bijections on all stars and costars. -/ +protected structure prefunctor.is_covering : Prop := +(star_bijective : ∀ u, bijective (φ.star u)) +(costar_bijective : ∀ u, bijective (φ.costar u)) + +@[simp] lemma prefunctor.is_covering.map_injective (hφ : φ.is_covering) {u v : U} : + injective (λ (f : u ⟶ v), φ.map f) := +begin + rintro f g he, + have : φ.star u (quiver.star.mk f) = φ.star u (quiver.star.mk g) := by simpa using he, + simpa using (hφ.star_bijective u).left this, +end + +lemma prefunctor.is_covering.comp (hφ : φ.is_covering) (hψ : ψ.is_covering) : + (φ ⋙q ψ).is_covering := +⟨λ u, (hψ.star_bijective _).comp (hφ.star_bijective _), + λ u, (hψ.costar_bijective _).comp (hφ.costar_bijective _)⟩ + +lemma prefunctor.is_covering.of_comp_right (hψ : ψ.is_covering) (hφψ : (φ ⋙q ψ).is_covering) : + φ.is_covering := +⟨λ u, (bijective.of_comp_iff' (hψ.star_bijective _) _).mp (hφψ.star_bijective _), + λ u, (bijective.of_comp_iff' (hψ.costar_bijective _) _).mp (hφψ.costar_bijective _)⟩ + +lemma prefunctor.is_covering.of_comp_left (hφ : φ.is_covering) (hφψ : (φ ⋙q ψ).is_covering) + (φsur : surjective φ.obj) : ψ.is_covering := +begin + refine ⟨λ v, _, λ v, _⟩; + obtain ⟨u, rfl⟩ := φsur v, + exacts [(bijective.of_comp_iff _ (hφ.star_bijective u)).mp (hφψ.star_bijective u), + (bijective.of_comp_iff _ (hφ.costar_bijective u)).mp (hφψ.costar_bijective u)], +end + +/-- The star of the symmetrification of a quiver at a vertex `u` is equivalent to the sum of the +star and the costar at `u` in the original quiver. -/ +def quiver.symmetrify_star (u : U) : + quiver.star (symmetrify.of.obj u) ≃ quiver.star u ⊕ quiver.costar u := +equiv.sigma_sum_distrib _ _ + +/-- The costar of the symmetrification of a quiver at a vertex `u` is equivalent to the sum of the +costar and the star at `u` in the original quiver. -/ +def quiver.symmetrify_costar (u : U) : + quiver.costar (symmetrify.of.obj u) ≃ quiver.costar u ⊕ quiver.star u := +equiv.sigma_sum_distrib _ _ + +lemma prefunctor.symmetrify_star (u : U) : + φ.symmetrify.star u = (quiver.symmetrify_star _).symm ∘ + sum.map (φ.star u) (φ.costar u) ∘ quiver.symmetrify_star u := +begin + rw equiv.eq_symm_comp, + ext ⟨v, (f|g)⟩; + simp [quiver.symmetrify_star], +end + +protected lemma prefunctor.symmetrify_costar (u : U) : + φ.symmetrify.costar u = (quiver.symmetrify_costar _).symm ∘ + sum.map (φ.costar u) (φ.star u) ∘ quiver.symmetrify_costar u := +begin + rw equiv.eq_symm_comp, + ext ⟨v, (f|g)⟩; + simp [quiver.symmetrify_costar], +end + +protected lemma prefunctor.is_covering.symmetrify (hφ : φ.is_covering) : φ.symmetrify.is_covering := +begin + refine ⟨λ u, _, λ u, _⟩; + simp [φ.symmetrify_star, φ.symmetrify_costar, hφ.star_bijective u, hφ.costar_bijective u], +end + +/-- The path star at a vertex `u` is the type of all paths starting at `u`. +The type `quiver.path_star u` is defined to be `Σ v : U, path u v`. -/ +@[reducible] def quiver.path_star (u : U) := Σ v : U, path u v + +/-- Constructor for `quiver.path_star`. Defined to be `sigma.mk`. -/ +@[reducible] protected def quiver.path_star.mk {u v : U} (p : path u v) : + quiver.path_star u := ⟨_, p⟩ + +/-- A prefunctor induces a map of path stars. -/ +def prefunctor.path_star (u : U) : quiver.path_star u → quiver.path_star (φ.obj u) := +λ p, quiver.path_star.mk (φ.map_path p.2) + +@[simp] lemma prefunctor.path_star_apply {u v : U} (p : path u v) : + φ.path_star u (quiver.path_star.mk p) = quiver.path_star.mk (φ.map_path p) := rfl + +theorem prefunctor.path_star_injective (hφ : ∀ u, injective (φ.star u)) (u : U) : + injective (φ.path_star u) := +begin + dsimp [prefunctor.path_star, quiver.path_star.mk], + rintro ⟨v₁, p₁⟩, + induction p₁ with x₁ y₁ p₁ e₁ ih; + rintro ⟨y₂, p₂⟩; cases p₂ with x₂ _ p₂ e₂; + intro h; + simp only [prefunctor.path_star_apply, prefunctor.map_path_nil, + prefunctor.map_path_cons] at h, + { exfalso, + cases h with h h', + rw [←path.eq_cast_iff_heq rfl h.symm, path.cast_cons] at h', + exact (path.nil_ne_cons _ _) h', }, + { exfalso, + cases h with h h', + rw [←path.cast_eq_iff_heq rfl h, path.cast_cons] at h', + exact (path.cons_ne_nil _ _) h', }, + { cases h with hφy h', + rw [←path.cast_eq_iff_heq rfl hφy, path.cast_cons, path.cast_rfl_rfl] at h', + have hφx := path.obj_eq_of_cons_eq_cons h', + have hφp := path.heq_of_cons_eq_cons h', + have hφe := heq.trans (hom.cast_heq rfl hφy _).symm (path.hom_heq_of_cons_eq_cons h'), + have h_path_star : φ.path_star u ⟨x₁, p₁⟩ = φ.path_star u ⟨x₂, p₂⟩, + { simp only [prefunctor.path_star_apply, sigma.mk.inj_iff], exact ⟨hφx, hφp⟩, }, + cases ih h_path_star, + have h_star : φ.star x₁ ⟨y₁, e₁⟩ = φ.star x₁ ⟨y₂, e₂⟩, + { simp only [prefunctor.star_apply, sigma.mk.inj_iff], exact ⟨hφy, hφe⟩, }, + cases hφ x₁ h_star, + refl, }, +end + +theorem prefunctor.path_star_surjective (hφ : ∀ u, surjective (φ.star u)) (u : U) : + surjective (φ.path_star u) := +begin + dsimp [prefunctor.path_star, quiver.path_star.mk], + rintro ⟨v, p⟩, + induction p with v' v'' p' ev ih, + { use ⟨u, path.nil⟩, + simp only [prefunctor.map_path_nil, eq_self_iff_true, heq_iff_eq, and_self], }, + { obtain ⟨⟨u', q'⟩, h⟩ := ih, + simp only at h, + obtain ⟨rfl,rfl⟩ := h, + obtain ⟨⟨u'', eu⟩, k⟩ := hφ u' ⟨_, ev⟩, + simp at k, + obtain ⟨rfl,rfl⟩ := k, + use ⟨_, q'.cons eu⟩, + simp only [prefunctor.map_path_cons, eq_self_iff_true, heq_iff_eq, and_self], } +end + +theorem prefunctor.path_star_bijective (hφ : ∀ u, bijective (φ.star u)) (u : U) : + bijective (φ.path_star u) := +⟨φ.path_star_injective (λ u, (hφ u).1) _, φ.path_star_surjective (λ u, (hφ u).2) _⟩ + +namespace prefunctor.is_covering +variable {φ} + +protected theorem path_star_bijective (hφ : φ.is_covering) (u : U) : + bijective (φ.path_star u) := φ.path_star_bijective hφ.1 u + +end prefunctor.is_covering + +section has_involutive_reverse + +variables [has_involutive_reverse U] [has_involutive_reverse V] [prefunctor.map_reverse φ] + +/-- In a quiver with involutive inverses, the star and costar at every vertex are equivalent. +This map is induced by `quiver.reverse`. -/ +@[simps] def quiver.star_equiv_costar (u : U) : + quiver.star u ≃ quiver.costar u := +{ to_fun := λ e, ⟨e.1, reverse e.2⟩, + inv_fun := λ e, ⟨e.1, reverse e.2⟩, + left_inv := λ e, by simp [sigma.ext_iff], + right_inv := λ e, by simp [sigma.ext_iff] } + +@[simp] lemma quiver.star_equiv_costar_apply {u v : U} (e : u ⟶ v) : + quiver.star_equiv_costar u (quiver.star.mk e) = quiver.costar.mk (reverse e) := rfl + +@[simp] lemma quiver.star_equiv_costar_symm_apply {u v : U} (e : u ⟶ v) : + (quiver.star_equiv_costar v).symm (quiver.costar.mk e) = quiver.star.mk (reverse e) := rfl + +lemma prefunctor.costar_conj_star (u : U) : + φ.costar u = + quiver.star_equiv_costar (φ.obj u) ∘ φ.star u ∘ (quiver.star_equiv_costar u).symm := +by { ext ⟨v, f⟩; simp, } + +lemma prefunctor.bijective_costar_iff_bijective_star (u : U) : + bijective (φ.costar u) ↔ bijective (φ.star u) := +begin + rw [prefunctor.costar_conj_star, bijective.of_comp_iff', bijective.of_comp_iff]; + exact equiv.bijective _, +end + +lemma prefunctor.is_covering_of_bijective_star (h : ∀ u, bijective (φ.star u)) : + φ.is_covering := ⟨h, λ u, (φ.bijective_costar_iff_bijective_star u).2 (h u)⟩ + +lemma prefunctor.is_covering_of_bijective_costar (h : ∀ u, bijective (φ.costar u)) : + φ.is_covering := ⟨λ u, (φ.bijective_costar_iff_bijective_star u).1 (h u), h⟩ + +end has_involutive_reverse From 0c1d80f5a86b36c1db32e021e8d19ae7809d5b79 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 28 Jul 2023 10:11:09 +0000 Subject: [PATCH 1275/1291] feat(linear_algebra/orientation): add `orientation.reindex` (#19236) --- src/linear_algebra/alternating.lean | 30 ++++++++++++++++++++++++ src/linear_algebra/determinant.lean | 5 ++++ src/linear_algebra/orientation.lean | 36 ++++++++++++++++++++++++++--- 3 files changed, 68 insertions(+), 3 deletions(-) diff --git a/src/linear_algebra/alternating.lean b/src/linear_algebra/alternating.lean index 0fa1eccdefb48..4f23645d2b111 100644 --- a/src/linear_algebra/alternating.lean +++ b/src/linear_algebra/alternating.lean @@ -581,6 +581,12 @@ rfl (f + g).dom_dom_congr σ = f.dom_dom_congr σ + g.dom_dom_congr σ := rfl +@[simp] lemma dom_dom_congr_smul {S : Type*} + [monoid S] [distrib_mul_action S N] [smul_comm_class R S N] (σ : ι ≃ ι') (c : S) + (f : alternating_map R M N ι) : + (c • f).dom_dom_congr σ = c • f.dom_dom_congr σ := +rfl + /-- `alternating_map.dom_dom_congr` as an equivalence. This is declared separately because it does not work with dot notation. -/ @@ -593,6 +599,30 @@ def dom_dom_congr_equiv (σ : ι ≃ ι') : right_inv := λ m, by { ext, simp [function.comp] }, map_add' := dom_dom_congr_add σ } +section dom_dom_lcongr +variables (S : Type*) [semiring S] [module S N] [smul_comm_class R S N] + +/-- `alternating_map.dom_dom_congr` as a linear equivalence. -/ +@[simps apply symm_apply] +def dom_dom_lcongr (σ : ι ≃ ι') : alternating_map R M N ι ≃ₗ[S] alternating_map R M N ι' := +{ to_fun := dom_dom_congr σ, + inv_fun := dom_dom_congr σ.symm, + left_inv := λ f, by { ext, simp [function.comp] }, + right_inv := λ m, by { ext, simp [function.comp] }, + map_add' := dom_dom_congr_add σ, + map_smul' := dom_dom_congr_smul σ } + +@[simp] lemma dom_dom_lcongr_refl : + (dom_dom_lcongr S (equiv.refl ι) : alternating_map R M N ι ≃ₗ[S] alternating_map R M N ι) = + linear_equiv.refl _ _ := +linear_equiv.ext dom_dom_congr_refl + +@[simp] lemma dom_dom_lcongr_to_add_equiv (σ : ι ≃ ι') : + (dom_dom_lcongr S σ : alternating_map R M N ι ≃ₗ[S] alternating_map R M N ι').to_add_equiv + = dom_dom_congr_equiv σ := rfl + +end dom_dom_lcongr + /-- The results of applying `dom_dom_congr` to two maps are equal if and only if those maps are. -/ @[simp] lemma dom_dom_congr_eq_iff (σ : ι ≃ ι') (f g : alternating_map R M N ι) : f.dom_dom_congr σ = g.dom_dom_congr σ ↔ f = g := diff --git a/src/linear_algebra/determinant.lean b/src/linear_algebra/determinant.lean index 80d805eccb482..c005163173ff9 100644 --- a/src/linear_algebra/determinant.lean +++ b/src/linear_algebra/determinant.lean @@ -559,6 +559,11 @@ lemma basis.det_reindex {ι' : Type*} [fintype ι'] [decidable_eq ι'] (b.reindex e).det v = b.det (v ∘ e) := by rw [basis.det_apply, basis.to_matrix_reindex', det_reindex_alg_equiv, basis.det_apply] +lemma basis.det_reindex' {ι' : Type*} [fintype ι'] [decidable_eq ι'] + (b : basis ι R M) (e : ι ≃ ι') : + (b.reindex e).det = b.det.dom_dom_congr e := +alternating_map.ext $ λ _, basis.det_reindex _ _ _ + lemma basis.det_reindex_symm {ι' : Type*} [fintype ι'] [decidable_eq ι'] (b : basis ι R M) (v : ι → M) (e : ι' ≃ ι) : (b.reindex e.symm).det (v ∘ e) = b.det v := diff --git a/src/linear_algebra/orientation.lean b/src/linear_algebra/orientation.lean index a17ec1101eba4..107bd5f34b839 100644 --- a/src/linear_algebra/orientation.lean +++ b/src/linear_algebra/orientation.lean @@ -43,7 +43,7 @@ section ordered_comm_semiring variables (R : Type*) [strict_ordered_comm_semiring R] variables (M : Type*) [add_comm_monoid M] [module R M] variables {N : Type*} [add_comm_monoid N] [module R N] -variables (ι : Type*) +variables (ι ι' : Type*) /-- An orientation of a module, intended to be used when `ι` is a `fintype` with the same cardinality as a basis. -/ @@ -73,6 +73,27 @@ by rw [orientation.map, alternating_map.dom_lcongr_refl, module.ray.map_refl] @[simp] lemma orientation.map_symm (e : M ≃ₗ[R] N) : (orientation.map ι e).symm = orientation.map ι e.symm := rfl +section reindex +variables (R M) {ι ι'} + +/-- An equivalence between indices implies an equivalence between orientations. -/ +def orientation.reindex (e : ι ≃ ι') : orientation R M ι ≃ orientation R M ι' := +module.ray.map $ alternating_map.dom_dom_lcongr R e + +@[simp] lemma orientation.reindex_apply (e : ι ≃ ι') (v : alternating_map R M R ι) + (hv : v ≠ 0) : + orientation.reindex R M e (ray_of_ne_zero _ v hv) = ray_of_ne_zero _ (v.dom_dom_congr e) + (mt (v.dom_dom_congr_eq_zero_iff e).mp hv) := rfl + +@[simp] lemma orientation.reindex_refl : + (orientation.reindex R M $ equiv.refl ι) = equiv.refl _ := +by rw [orientation.reindex, alternating_map.dom_dom_lcongr_refl, module.ray.map_refl] + +@[simp] lemma orientation.reindex_symm (e : ι ≃ ι') : + (orientation.reindex R M e).symm = orientation.reindex R M e.symm := rfl + +end reindex + /-- A module is canonically oriented with respect to an empty index type. -/ @[priority 100] instance is_empty.oriented [nontrivial R] [is_empty ι] : module.oriented R M ι := @@ -107,9 +128,14 @@ variables {M N : Type*} [add_comm_group M] [add_comm_group N] [module R M] [modu orientation.map ι f (-x) = - orientation.map ι f x := module.ray.map_neg _ x +@[simp] protected lemma orientation.reindex_neg {ι ι' : Type*} (e : ι ≃ ι') + (x : orientation R M ι) : + orientation.reindex R M e (-x) = - orientation.reindex R M e x := +module.ray.map_neg _ x + namespace basis -variables {ι : Type*} +variables {ι ι' : Type*} /-- The value of `orientation.map` when the index type has the cardinality of a basis, in terms of `f.det`. -/ @@ -125,7 +151,7 @@ begin basis.det_self, mul_one, smul_eq_mul, mul_comm, mul_smul, linear_equiv.coe_inv_det], end -variables [fintype ι] [decidable_eq ι] +variables [fintype ι] [decidable_eq ι] [fintype ι'] [decidable_eq ι'] /-- The orientation given by a basis. -/ protected def orientation [nontrivial R] (e : basis ι R M) : orientation R M ι := @@ -135,6 +161,10 @@ lemma orientation_map [nontrivial R] (e : basis ι R M) (f : M ≃ₗ[R] N) : (e.map f).orientation = orientation.map ι f e.orientation := by simp_rw [basis.orientation, orientation.map_apply, basis.det_map'] +lemma orientation_reindex [nontrivial R] (e : basis ι R M) + (eι : ι ≃ ι') : (e.reindex eι).orientation = orientation.reindex R M eι e.orientation := +by simp_rw [basis.orientation, orientation.reindex_apply, basis.det_reindex'] + /-- The orientation given by a basis derived using `units_smul`, in terms of the product of those units. -/ lemma orientation_units_smul [nontrivial R] (e : basis ι R M) (w : ι → units R) : From 3b52265189f3fb43aa631edffce5d060fafaf82f Mon Sep 17 00:00:00 2001 From: AlexKontorovich <25316162+hrmacbeth@users.noreply.github.com> Date: Fri, 28 Jul 2023 12:48:03 +0000 Subject: [PATCH 1276/1291] feat(measure_theory/measure/haar_quotient): the Unfolding Trick (#18863) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We prove the "unfolding trick": Given a subgroup `Γ` of a group `G`, the integral of a function `f` on `G` times the lift to `G` of a function `g` on the coset space `G ⧸ Γ` with respect to a right-invariant measure `μ` on `G`, is equal to the integral over the coset space of the automorphization of `f` times `g`. We also prove the following simplified version: Given a subgroup `Γ` of a group `G`, the integral of a function `f` on `G` with respect to a right-invariant measure `μ` is equal to the integral over the coset space `G ⧸ Γ` of the automorphization of `f`. A question: is it possible to deduce `ae_strongly_measurable (quotient_group.automorphize f) μ_𝓕` from `ae_strongly_measurable f μ` (as opposed to assuming it as a hypothesis in the main theorem)? It seems quite plausible... Co-authored-by: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Co-authored-by: Alex Kontorovich <58564076+AlexKontorovich@users.noreply.github.com> Co-authored-by: AlexKontorovich <58564076+AlexKontorovich@users.noreply.github.com> --- src/group_theory/group_action/group.lean | 11 ++ .../function/strongly_measurable/basic.lean | 7 + .../group/fundamental_domain.lean | 9 + src/measure_theory/measure/haar/quotient.lean | 183 +++++++++++++++++- src/topology/algebra/infinite_sum/basic.lean | 141 ++++++++++++++ 5 files changed, 343 insertions(+), 8 deletions(-) diff --git a/src/group_theory/group_action/group.lean b/src/group_theory/group_action/group.lean index 02971812f5784..be6086db460c1 100644 --- a/src/group_theory/group_action/group.lean +++ b/src/group_theory/group_action/group.lean @@ -191,6 +191,17 @@ def distrib_mul_action.to_add_equiv (x : α) : β ≃+ β := { .. distrib_mul_action.to_add_monoid_hom β x, .. mul_action.to_perm_hom α β x } +/-- Each non-zero element of a `group_with_zero` defines an additive monoid isomorphism of an +`add_monoid` on which it acts distributively. + +This is a stronger version of `distrib_mul_action.to_add_monoid_hom`. -/ +def distrib_mul_action.to_add_equiv₀ {α : Type*} (β : Type*) [group_with_zero α] [add_monoid β] + [distrib_mul_action α β] (x : α) (hx : x ≠ 0) : β ≃+ β := +{ inv_fun := λ b, x⁻¹ • b, + left_inv := inv_smul_smul₀ hx, + right_inv := smul_inv_smul₀ hx, + .. distrib_mul_action.to_add_monoid_hom β x, } + variables (α β) /-- Each element of the group defines an additive monoid isomorphism. diff --git a/src/measure_theory/function/strongly_measurable/basic.lean b/src/measure_theory/function/strongly_measurable/basic.lean index c7220c3cb2d68..a2efb0496676c 100644 --- a/src/measure_theory/function/strongly_measurable/basic.lean +++ b/src/measure_theory/function/strongly_measurable/basic.lean @@ -1740,6 +1740,13 @@ end end ae_strongly_measurable +lemma ae_strongly_measurable_of_absolutely_continuous {α β : Type*} [measurable_space α] + [topological_space β] {μ ν : measure α} (h : ν ≪ μ) (g : α → β) + (hμ : ae_strongly_measurable g μ) : ae_strongly_measurable g ν := +begin + obtain ⟨g₁, hg₁, hg₁'⟩ := hμ, + refine ⟨g₁, hg₁, h.ae_eq hg₁'⟩, +end /-! ## Almost everywhere finitely strongly measurable functions -/ diff --git a/src/measure_theory/group/fundamental_domain.lean b/src/measure_theory/group/fundamental_domain.lean index ec47f9ce9daf4..c97dba8abe1bc 100644 --- a/src/measure_theory/group/fundamental_domain.lean +++ b/src/measure_theory/group/fundamental_domain.lean @@ -212,6 +212,11 @@ calc ∫⁻ x, f x ∂μ = ∑' g : G, ∫⁻ x in g • s, f x ∂μ : h.linteg tsum_congr $ λ g, ((measure_preserving_smul g⁻¹ μ).set_lintegral_comp_emb (measurable_embedding_const_smul _) _ _).symm + +@[to_additive] lemma lintegral_eq_tsum'' (h : is_fundamental_domain G s μ) (f : α → ℝ≥0∞) : + ∫⁻ x, f x ∂μ = ∑' g : G, ∫⁻ x in s, f (g • x) ∂μ := +(lintegral_eq_tsum' h f).trans ((equiv.inv G).tsum_eq (λ g, ∫⁻ (x : α) in s, f (g • x) ∂μ)) + @[to_additive] lemma set_lintegral_eq_tsum (h : is_fundamental_domain G s μ) (f : α → ℝ≥0∞) (t : set α) : ∫⁻ x in t, f x ∂μ = ∑' g : G, ∫⁻ x in t ∩ g • s, f x ∂μ := calc ∫⁻ x in t, f x ∂μ = ∑' g : G, ∫⁻ x in g • s, f x ∂(μ.restrict t) : @@ -357,6 +362,10 @@ calc ∫ x, f x ∂μ = ∑' g : G, ∫ x in g • s, f x ∂μ : h.integral_eq_ tsum_congr $ λ g, (measure_preserving_smul g⁻¹ μ).set_integral_image_emb (measurable_embedding_const_smul _) _ _ +@[to_additive] lemma integral_eq_tsum'' (h : is_fundamental_domain G s μ) + (f : α → E) (hf : integrable f μ) : ∫ x, f x ∂μ = ∑' g : G, ∫ x in s, f (g • x) ∂μ := +(integral_eq_tsum' h f hf).trans ((equiv.inv G).tsum_eq (λ g, ∫ (x : α) in s, f (g • x) ∂μ)) + @[to_additive] lemma set_integral_eq_tsum (h : is_fundamental_domain G s μ) {f : α → E} {t : set α} (hf : integrable_on f t μ) : ∫ x in t, f x ∂μ = ∑' g : G, ∫ x in t ∩ g • s, f x ∂μ := diff --git a/src/measure_theory/measure/haar/quotient.lean b/src/measure_theory/measure/haar/quotient.lean index 03f20400f0398..5bd1067dbc4e5 100644 --- a/src/measure_theory/measure/haar/quotient.lean +++ b/src/measure_theory/measure/haar/quotient.lean @@ -33,8 +33,10 @@ Note that a group `G` with Haar measure that is both left and right invariant is **unimodular**. -/ -open set measure_theory topological_space measure_theory.measure -open_locale pointwise nnreal +noncomputable theory + +open set measure_theory topological_space measure_theory.measure quotient_group +open_locale pointwise measure_theory topology big_operators nnreal ennreal variables {G : Type*} [group G] [measurable_space G] [topological_space G] [topological_group G] [borel_space G] @@ -116,8 +118,6 @@ lemma measure_theory.is_fundamental_domain.is_mul_left_invariant_map [subgroup.n { exact hA, }, end } -variables [t2_space (G ⧸ Γ)] [second_countable_topology (G ⧸ Γ)] (K : positive_compacts (G ⧸ Γ)) - /-- Given a normal subgroup `Γ` of a topological group `G` with Haar measure `μ`, which is also right-invariant, and a finite volume fundamental domain `𝓕`, the pushforward to the quotient group `G ⧸ Γ` of the restriction of `μ` to `𝓕` is a multiple of Haar measure on `G ⧸ Γ`. -/ @@ -125,7 +125,8 @@ variables [t2_space (G ⧸ Γ)] [second_countable_topology (G ⧸ Γ)] (K : posi `μ`, which is also right-invariant, and a finite volume fundamental domain `𝓕`, the pushforward to the quotient group `G ⧸ Γ` of the restriction of `μ` to `𝓕` is a multiple of Haar measure on `G ⧸ Γ`."] -lemma measure_theory.is_fundamental_domain.map_restrict_quotient [subgroup.normal Γ] +lemma measure_theory.is_fundamental_domain.map_restrict_quotient [t2_space (G ⧸ Γ)] + [second_countable_topology (G ⧸ Γ)] (K : positive_compacts (G ⧸ Γ)) [subgroup.normal Γ] [measure_theory.measure.is_haar_measure μ] [μ.is_mul_right_invariant] (h𝓕_finite : μ 𝓕 < ⊤) : measure.map (quotient_group.mk' Γ) (μ.restrict 𝓕) = (μ (𝓕 ∩ (quotient_group.mk' Γ) ⁻¹' K)) • (measure_theory.measure.haar_measure K) := @@ -151,12 +152,178 @@ end topological group `G` with Haar measure `μ`, which is also right-invariant, and a finite volume fundamental domain `𝓕`, the quotient map to `G ⧸ Γ` is measure-preserving between appropriate multiples of Haar measure on `G` and `G ⧸ Γ`."] -lemma measure_preserving_quotient_group.mk' [subgroup.normal Γ] - [measure_theory.measure.is_haar_measure μ] [μ.is_mul_right_invariant] - (h𝓕_finite : μ 𝓕 < ⊤) (c : ℝ≥0) (h : μ (𝓕 ∩ (quotient_group.mk' Γ) ⁻¹' K) = c) : +lemma measure_preserving_quotient_group.mk' [t2_space (G ⧸ Γ)] [second_countable_topology (G ⧸ Γ)] + (K : positive_compacts (G ⧸ Γ)) [subgroup.normal Γ] [measure_theory.measure.is_haar_measure μ] + [μ.is_mul_right_invariant] (h𝓕_finite : μ 𝓕 < ⊤) (c : ℝ≥0) + (h : μ (𝓕 ∩ (quotient_group.mk' Γ) ⁻¹' K) = c) : measure_preserving (quotient_group.mk' Γ) (μ.restrict 𝓕) (c • (measure_theory.measure.haar_measure K)) := { measurable := continuous_quotient_mk.measurable, map_eq := by rw [h𝓕.map_restrict_quotient K h𝓕_finite, h]; refl } + +section + +local notation `μ_𝓕` := measure.map (@quotient_group.mk G _ Γ) (μ.restrict 𝓕) + +/-- The `ess_sup` of a function `g` on the quotient space `G ⧸ Γ` with respect to the pushforward + of the restriction, `μ_𝓕`, of a right-invariant measure `μ` to a fundamental domain `𝓕`, is the + same as the `ess_sup` of `g`'s lift to the universal cover `G` with respect to `μ`. -/ +@[to_additive "The `ess_sup` of a function `g` on the additive quotient space `G ⧸ Γ` with respect + to the pushforward of the restriction, `μ_𝓕`, of a right-invariant measure `μ` to a fundamental + domain `𝓕`, is the same as the `ess_sup` of `g`'s lift to the universal cover `G` with respect + to `μ`."] +lemma ess_sup_comp_quotient_group_mk [μ.is_mul_right_invariant] {g : G ⧸ Γ → ℝ≥0∞} + (g_ae_measurable : ae_measurable g μ_𝓕) : + ess_sup g μ_𝓕 = ess_sup (λ (x : G), g x) μ := +begin + have hπ : measurable (quotient_group.mk : G → G ⧸ Γ) := continuous_quotient_mk.measurable, + rw ess_sup_map_measure g_ae_measurable hπ.ae_measurable, + refine h𝓕.ess_sup_measure_restrict _, + rintros ⟨γ, hγ⟩ x, + dsimp, + congr' 1, + exact quotient_group.mk_mul_of_mem x hγ, +end + +/-- Given a quotient space `G ⧸ Γ` where `Γ` is `countable`, and the restriction, + `μ_𝓕`, of a right-invariant measure `μ` on `G` to a fundamental domain `𝓕`, a set + in the quotient which has `μ_𝓕`-measure zero, also has measure zero under the + folding of `μ` under the quotient. Note that, if `Γ` is infinite, then the folded map + will take the value `∞` on any open set in the quotient! -/ +@[to_additive "Given an additive quotient space `G ⧸ Γ` where `Γ` is `countable`, and the + restriction, `μ_𝓕`, of a right-invariant measure `μ` on `G` to a fundamental domain `𝓕`, a set + in the quotient which has `μ_𝓕`-measure zero, also has measure zero under the + folding of `μ` under the quotient. Note that, if `Γ` is infinite, then the folded map + will take the value `∞` on any open set in the quotient!"] +lemma _root_.measure_theory.is_fundamental_domain.absolutely_continuous_map + [μ.is_mul_right_invariant] : + map (quotient_group.mk : G → G ⧸ Γ) μ ≪ map (quotient_group.mk : G → G ⧸ Γ) (μ.restrict 𝓕) := +begin + set π : G → G ⧸ Γ := quotient_group.mk, + have meas_π : measurable π := continuous_quotient_mk.measurable, + apply absolutely_continuous.mk, + intros s s_meas hs, + rw map_apply meas_π s_meas at hs ⊢, + rw measure.restrict_apply at hs, + apply h𝓕.measure_zero_of_invariant _ _ hs, + { intros γ, + ext g, + rw [set.mem_smul_set_iff_inv_smul_mem, mem_preimage, mem_preimage], + congrm _ ∈ s, + convert quotient_group.mk_mul_of_mem g (γ⁻¹).2, }, + exact measurable_set_preimage meas_π s_meas, +end + +local attribute [-instance] quotient.measurable_space + +/-- This is a simple version of the **Unfolding Trick**: Given a subgroup `Γ` of a group `G`, the + integral of a function `f` on `G` with respect to a right-invariant measure `μ` is equal to the + integral over the quotient `G ⧸ Γ` of the automorphization of `f`. -/ +@[to_additive "This is a simple version of the **Unfolding Trick**: Given a subgroup `Γ` of an + additive group `G`, the integral of a function `f` on `G` with respect to a right-invariant + measure `μ` is equal to the integral over the quotient `G ⧸ Γ` of the automorphization of `f`."] +lemma quotient_group.integral_eq_integral_automorphize {E : Type*} [normed_add_comm_group E] + [complete_space E] [normed_space ℝ E] [μ.is_mul_right_invariant] {f : G → E} + (hf₁ : integrable f μ) (hf₂ : ae_strongly_measurable (automorphize f) μ_𝓕) : + ∫ x : G, f x ∂μ = ∫ x : G ⧸ Γ, automorphize f x ∂μ_𝓕 := +calc ∫ x : G, f x ∂μ = ∑' γ : Γ.opposite, ∫ x in 𝓕, f (γ • x) ∂μ : h𝓕.integral_eq_tsum'' f hf₁ +... = ∫ x in 𝓕, ∑' γ : Γ.opposite, f (γ • x) ∂μ : + begin + rw integral_tsum, + { exact λ i, (hf₁.1.comp_quasi_measure_preserving + (measure_preserving_smul i μ).quasi_measure_preserving).restrict, }, + { rw ← h𝓕.lintegral_eq_tsum'' (λ x, ‖f x‖₊), + exact ne_of_lt hf₁.2, }, + end +... = ∫ x : G ⧸ Γ, automorphize f x ∂μ_𝓕 : + (integral_map continuous_quotient_mk.ae_measurable hf₂).symm + +/-- This is the **Unfolding Trick**: Given a subgroup `Γ` of a group `G`, the integral of a + function `f` on `G` times the lift to `G` of a function `g` on the quotient `G ⧸ Γ` with respect + to a right-invariant measure `μ` on `G`, is equal to the integral over the quotient of the + automorphization of `f` times `g`. -/ +lemma quotient_group.integral_mul_eq_integral_automorphize_mul {K : Type*} [normed_field K] + [complete_space K] [normed_space ℝ K] [μ.is_mul_right_invariant] {f : G → K} + (f_ℒ_1 : integrable f μ) {g : G ⧸ Γ → K} (hg : ae_strongly_measurable g μ_𝓕) + (g_ℒ_infinity : ess_sup (λ x, ↑‖g x‖₊) μ_𝓕 ≠ ∞) + (F_ae_measurable : ae_strongly_measurable (quotient_group.automorphize f) μ_𝓕) : + ∫ x : G, g (x : G ⧸ Γ) * (f x) ∂μ = ∫ x : G ⧸ Γ, g x * (quotient_group.automorphize f x) ∂μ_𝓕 := +begin + let π : G → G ⧸ Γ := quotient_group.mk, + have H₀ : quotient_group.automorphize ((g ∘ π) * f) = g * (quotient_group.automorphize f) := + quotient_group.automorphize_smul_left f g, + calc ∫ (x : G), g (π x) * f x ∂μ = + ∫ (x : G ⧸ Γ), quotient_group.automorphize ((g ∘ π) * f) x ∂μ_𝓕 : _ + ... = ∫ (x : G ⧸ Γ), g x * (quotient_group.automorphize f x) ∂μ_𝓕 : by simp [H₀], + have meas_π : measurable π := continuous_quotient_mk.measurable, + have H₁ : integrable ((g ∘ π) * f) μ, + { have : ae_strongly_measurable (λ x : G, g (x : G ⧸ Γ)) μ, + { refine (ae_strongly_measurable_of_absolutely_continuous _ _ hg).comp_measurable meas_π, + exact h𝓕.absolutely_continuous_map }, + refine integrable.ess_sup_smul f_ℒ_1 this _, + { have hg' : ae_strongly_measurable (λ x, ↑‖g x‖₊) μ_𝓕 := + (ennreal.continuous_coe.comp continuous_nnnorm).comp_ae_strongly_measurable hg, + rw [← ess_sup_comp_quotient_group_mk h𝓕 hg'.ae_measurable], + exact g_ℒ_infinity } }, + have H₂ : ae_strongly_measurable (quotient_group.automorphize ((g ∘ π) * f)) μ_𝓕, + { simp_rw [H₀], + exact hg.mul F_ae_measurable }, + apply quotient_group.integral_eq_integral_automorphize h𝓕 H₁ H₂, +end + +end + +section + +variables {G' : Type*} [add_group G'] [measurable_space G'] [topological_space G'] + [topological_add_group G'] [borel_space G'] + {μ' : measure G'} + {Γ' : add_subgroup G'} + [countable Γ'] [measurable_space (G' ⧸ Γ')] [borel_space (G' ⧸ Γ')] + {𝓕' : set G'} + +local notation `μ_𝓕` := measure.map (@quotient_add_group.mk G' _ Γ') (μ'.restrict 𝓕') + +/-- This is the **Unfolding Trick**: Given an additive subgroup `Γ'` of an additive group `G'`, the + integral of a function `f` on `G'` times the lift to `G'` of a function `g` on the quotient + `G' ⧸ Γ'` with respect to a right-invariant measure `μ` on `G'`, is equal to the integral over + the quotient of the automorphization of `f` times `g`. -/ +lemma quotient_add_group.integral_mul_eq_integral_automorphize_mul +{K : Type*} [normed_field K] + [complete_space K] [normed_space ℝ K] [μ'.is_add_right_invariant] {f : G' → K} + (f_ℒ_1 : integrable f μ') {g : G' ⧸ Γ' → K} (hg : ae_strongly_measurable g μ_𝓕) + (g_ℒ_infinity : ess_sup (λ x, ↑‖g x‖₊) μ_𝓕 ≠ ∞) + (F_ae_measurable : ae_strongly_measurable (quotient_add_group.automorphize f) μ_𝓕) + (h𝓕 : is_add_fundamental_domain Γ'.opposite 𝓕' μ') : + ∫ x : G', g (x : G' ⧸ Γ') * (f x) ∂μ' + = ∫ x : G' ⧸ Γ', g x * (quotient_add_group.automorphize f x) ∂μ_𝓕 := +begin + let π : G' → G' ⧸ Γ' := quotient_add_group.mk, + have H₀ : quotient_add_group.automorphize ((g ∘ π) * f) + = g * (quotient_add_group.automorphize f) := + quotient_add_group.automorphize_smul_left f g, + calc ∫ (x : G'), g (π x) * f x ∂μ' = + ∫ (x : G' ⧸ Γ'), quotient_add_group.automorphize ((g ∘ π) * f) x ∂μ_𝓕 : _ + ... = ∫ (x : G' ⧸ Γ'), g x * (quotient_add_group.automorphize f x) ∂μ_𝓕 : by simp [H₀], + have meas_π : measurable π := continuous_quotient_mk.measurable, + have H₁ : integrable ((g ∘ π) * f) μ', + { have : ae_strongly_measurable (λ x : G', g (x : G' ⧸ Γ')) μ', + { refine (ae_strongly_measurable_of_absolutely_continuous _ _ hg).comp_measurable meas_π, + exact h𝓕.absolutely_continuous_map }, + refine integrable.ess_sup_smul f_ℒ_1 this _, + { have hg' : ae_strongly_measurable (λ x, ↑‖g x‖₊) μ_𝓕 := + (ennreal.continuous_coe.comp continuous_nnnorm).comp_ae_strongly_measurable hg, + rw [← ess_sup_comp_quotient_add_group_mk h𝓕 hg'.ae_measurable], + exact g_ℒ_infinity } }, + have H₂ : ae_strongly_measurable (quotient_add_group.automorphize ((g ∘ π) * f)) μ_𝓕, + { simp_rw [H₀], + exact hg.mul F_ae_measurable }, + apply quotient_add_group.integral_eq_integral_automorphize h𝓕 H₁ H₂, +end + +end + +attribute [to_additive quotient_group.integral_mul_eq_integral_automorphize_mul] + quotient_add_group.integral_mul_eq_integral_automorphize_mul diff --git a/src/topology/algebra/infinite_sum/basic.lean b/src/topology/algebra/infinite_sum/basic.lean index 9d0486d2144d3..69334b3838d57 100644 --- a/src/topology/algebra/infinite_sum/basic.lean +++ b/src/topology/algebra/infinite_sum/basic.lean @@ -1147,9 +1147,46 @@ hf.map (distrib_mul_action.to_add_monoid_hom α _) $ continuous_const_smul _ lemma summable.const_smul (b : γ) (hf : summable f) : summable (λ i, b • f i) := (hf.has_sum.const_smul _).summable +/-- Infinite sums commute with scalar multiplication. Version for scalars living in a `monoid`, but + requiring a summability hypothesis. -/ lemma tsum_const_smul [t2_space α] (b : γ) (hf : summable f) : ∑' i, b • f i = b • ∑' i, f i := (hf.has_sum.const_smul _).tsum_eq +/-- Infinite sums commute with scalar multiplication. Version for scalars living in a `group`, but + not requiring any summability hypothesis. -/ +lemma tsum_const_smul' {γ : Type*} [group γ] [distrib_mul_action γ α] + [has_continuous_const_smul γ α] [t2_space α] (g : γ) : + ∑' (i : β), g • f i = g • ∑' (i : β), f i := +begin + by_cases hf : summable f, + { exact tsum_const_smul _ hf, }, + rw tsum_eq_zero_of_not_summable hf, + simp only [smul_zero], + let mul_g : α ≃+ α := distrib_mul_action.to_add_equiv α g, + apply tsum_eq_zero_of_not_summable, + change ¬ summable (mul_g ∘ f), + rwa summable.map_iff_of_equiv mul_g; apply continuous_const_smul, +end + +/-- Infinite sums commute with scalar multiplication. Version for scalars living in a + `division_ring`; no summability hypothesis. This could be made to work for a + `[group_with_zero γ]` if there was such a thing as `distrib_mul_action_with_zero`. -/ +lemma tsum_const_smul'' {γ : Type*} [division_ring γ] [module γ α] [has_continuous_const_smul γ α] + [t2_space α] (g : γ) : + ∑' (i : β), g • f i = g • ∑' (i : β), f i := +begin + by_cases hf : summable f, + { exact tsum_const_smul _ hf, }, + rw tsum_eq_zero_of_not_summable hf, + simp only [smul_zero], + by_cases hg : g = 0, + { simp [hg], }, + let mul_g : α ≃+ α := distrib_mul_action.to_add_equiv₀ α g hg, + apply tsum_eq_zero_of_not_summable, + change ¬ summable (mul_g ∘ f), + rwa summable.map_iff_of_equiv mul_g; apply continuous_const_smul, +end + end const_smul /-! ### Product and pi types -/ @@ -1256,3 +1293,107 @@ begin end end has_continuous_star + +section automorphize + +variables {M : Type*} [topological_space M] [add_comm_monoid M] [t2_space M] {R : Type*} + [division_ring R] [module R M] [has_continuous_const_smul R M] + +/-- Given a group `α` acting on a type `β`, and a function `f : β → M`, we "automorphize" `f` to a + function `β ⧸ α → M` by summing over `α` orbits, `b ↦ ∑' (a : α), f(a • b)`. -/ +@[to_additive "Given an additive group `α` acting on a type `β`, and a function `f : β → M`, + we automorphize `f` to a function `β ⧸ α → M` by summing over `α` orbits, + `b ↦ ∑' (a : α), f(a • b)`."] +def mul_action.automorphize [group α] [mul_action α β] (f : β → M) : + quotient (mul_action.orbit_rel α β) → M := +@quotient.lift _ _ (mul_action.orbit_rel α β) (λ b, ∑' (a : α), f(a • b)) +begin + rintros b₁ b₂ ⟨a, (rfl : a • b₂ = b₁)⟩, + simpa [mul_smul] using (equiv.mul_right a).tsum_eq (λ a', f (a' • b₂)), +end + +/-- Automorphization of a function into an `R`-`module` distributes, that is, commutes with the + `R`-scalar multiplication. -/ +lemma mul_action.automorphize_smul_left [group α] [mul_action α β] (f : β → M) + (g : quotient (mul_action.orbit_rel α β) → R) : + mul_action.automorphize ((g ∘ quotient.mk') • f) + = g • (mul_action.automorphize f : quotient (mul_action.orbit_rel α β) → M) := +begin + ext x, + apply quotient.induction_on' x, + intro b, + simp only [mul_action.automorphize, pi.smul_apply', function.comp_app], + set π : β → quotient (mul_action.orbit_rel α β) := quotient.mk', + have H₁ : ∀ a : α, π (a • b) = π b, + { intro a, + rw quotient.eq_rel, + fconstructor, + exact a, + simp, }, + change ∑' a : α, g (π (a • b)) • f (a • b) = g (π b) • ∑' a : α, f (a • b), + simp_rw [H₁], + exact tsum_const_smul'' _, +end + +/-- Automorphization of a function into an `R`-`module` distributes, that is, commutes with the + `R`-scalar multiplication. -/ +lemma add_action.automorphize_smul_left [add_group α] [add_action α β] (f : β → M) + (g : quotient (add_action.orbit_rel α β) → R) : + add_action.automorphize ((g ∘ quotient.mk') • f) + = g • (add_action.automorphize f : quotient (add_action.orbit_rel α β) → M) := +begin + ext x, + apply quotient.induction_on' x, + intro b, + simp only [add_action.automorphize, pi.smul_apply', function.comp_app], + set π : β → quotient (add_action.orbit_rel α β) := quotient.mk', + have H₁ : ∀ a : α, π (a +ᵥ b) = π b, + { intro a, + rw quotient.eq_rel, + fconstructor, + exact a, + simp, }, + change ∑' a : α, g (π (a +ᵥ b)) • f (a +ᵥ b) = g (π b) • ∑' a : α, f (a +ᵥ b), + simp_rw [H₁], + exact tsum_const_smul'' _, +end + +attribute [to_additive mul_action.automorphize_smul_left] add_action.automorphize_smul_left + +section + +variables {G : Type*} [group G] {Γ : subgroup G} + +/-- Given a subgroup `Γ` of a group `G`, and a function `f : G → M`, we "automorphize" `f` to a + function `G ⧸ Γ → M` by summing over `Γ` orbits, `g ↦ ∑' (γ : Γ), f(γ • g)`. -/ +@[to_additive "Given a subgroup `Γ` of an additive group `G`, and a function `f : G → M`, we + automorphize `f` to a function `G ⧸ Γ → M` by summing over `Γ` orbits, + `g ↦ ∑' (γ : Γ), f(γ • g)`."] +def quotient_group.automorphize (f : G → M) : G ⧸ Γ → M := mul_action.automorphize f + +/-- Automorphization of a function into an `R`-`module` distributes, that is, commutes with the + `R`-scalar multiplication. -/ +lemma quotient_group.automorphize_smul_left (f : G → M) (g : G ⧸ Γ → R) : + quotient_group.automorphize ((g ∘ quotient.mk') • f) + = g • (quotient_group.automorphize f : G ⧸ Γ → M) := +mul_action.automorphize_smul_left f g + +end + +section + +variables {G : Type*} [add_group G] {Γ : add_subgroup G} + +/-- Automorphization of a function into an `R`-`module` distributes, that is, commutes with the `R` + -scalar multiplication. -/ +lemma quotient_add_group.automorphize_smul_left (f : G → M) (g : G ⧸ Γ → R) : + quotient_add_group.automorphize ((g ∘ quotient.mk') • f) + = g • (quotient_add_group.automorphize f : G ⧸ Γ → M) := +add_action.automorphize_smul_left f g + +end + +attribute [to_additive quotient_group.automorphize_smul_left] + quotient_add_group.automorphize_smul_left + +end automorphize From 63721b2c3eba6c325ecf8ae8cca27155a4f6306f Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Sun, 30 Jul 2023 10:18:54 +0000 Subject: [PATCH 1277/1291] feat(algebraic_topology/dold_kan): The Dold-Kan equivalence for pseudoabelian categories (#17925) --- .../dold_kan/equivalence_pseudoabelian.lean | 123 ++++++++++++++++++ src/category_theory/functor/category.lean | 3 + 2 files changed, 126 insertions(+) create mode 100644 src/algebraic_topology/dold_kan/equivalence_pseudoabelian.lean diff --git a/src/algebraic_topology/dold_kan/equivalence_pseudoabelian.lean b/src/algebraic_topology/dold_kan/equivalence_pseudoabelian.lean new file mode 100644 index 0000000000000..dd766e8a5e41f --- /dev/null +++ b/src/algebraic_topology/dold_kan/equivalence_pseudoabelian.lean @@ -0,0 +1,123 @@ +/- +Copyright (c) 2022 Joël Riou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joël Riou +-/ + +import algebraic_topology.dold_kan.equivalence_additive +import algebraic_topology.dold_kan.compatibility +import category_theory.idempotents.simplicial_object + +/-! + +# The Dold-Kan correspondence for pseudoabelian categories + +In this file, for any idempotent complete additive category `C`, +the Dold-Kan equivalence +`idempotents.dold_kan.equivalence C : simplicial_object C ≌ chain_complex C ℕ` +is obtained. It is deduced from the equivalence +`preadditive.dold_kan.equivalence` between the respective idempotent +completions of these categories using the fact that when `C` is idempotent complete, +then both `simplicial_object C` and `chain_complex C ℕ` are idempotent complete. + +The construction of `idempotents.dold_kan.equivalence` uses the tools +introduced in the file `compatibility.lean`. Doing so, the functor +`idempotents.dold_kan.N` of the equivalence is +the composition of `N₁ : simplicial_object C ⥤ karoubi (chain_complex C ℕ)` +(defined in `functor_n.lean`) and the inverse of the equivalence +`chain_complex C ℕ ≌ karoubi (chain_complex C ℕ)`. The functor +`idempotents.dold_kan.Γ` of the equivalence is by definition the functor +`Γ₀` introduced in `functor_gamma.lean`. + +-/ + +noncomputable theory + +open category_theory category_theory.category category_theory.limits category_theory.idempotents + +variables {C : Type*} [category C] [preadditive C] [is_idempotent_complete C] + [has_finite_coproducts C] + +namespace category_theory + +namespace idempotents + +namespace dold_kan + +open algebraic_topology.dold_kan + +/-- The functor `N` for the equivalence is obtained by composing +`N' : simplicial_object C ⥤ karoubi (chain_complex C ℕ)` and the inverse +of the equivalence `chain_complex C ℕ ≌ karoubi (chain_complex C ℕ)`. -/ +@[simps, nolint unused_arguments] +def N : simplicial_object C ⥤ chain_complex C ℕ := +N₁ ⋙ (to_karoubi_equivalence _).inverse + +/-- The functor `Γ` for the equivalence is `Γ'`. -/ +@[simps, nolint unused_arguments] +def Γ : chain_complex C ℕ ⥤ simplicial_object C := Γ₀ + +lemma hN₁ : (to_karoubi_equivalence (simplicial_object C)).functor ⋙ + preadditive.dold_kan.equivalence.functor = N₁ := +functor.congr_obj (functor_extension₁_comp_whiskering_left_to_karoubi _ _) N₁ + +lemma hΓ₀ : (to_karoubi_equivalence (chain_complex C ℕ)).functor ⋙ + preadditive.dold_kan.equivalence.inverse = Γ ⋙ (to_karoubi_equivalence _).functor := +functor.congr_obj (functor_extension₂_comp_whiskering_left_to_karoubi _ _) Γ₀ + +/-- The Dold-Kan equivalence for pseudoabelian categories given +by the functors `N` and `Γ`. It is obtained by applying the results in +`compatibility.lean` to the equivalence `preadditive.dold_kan.equivalence`. -/ +def equivalence : simplicial_object C ≌ chain_complex C ℕ := +compatibility.equivalence (eq_to_iso hN₁) (eq_to_iso hΓ₀) + +lemma equivalence_functor : (equivalence : simplicial_object C ≌ _).functor = N := rfl +lemma equivalence_inverse : (equivalence : simplicial_object C ≌ _).inverse = Γ := rfl + +/-- The natural isomorphism `NΓ' satisfies the compatibility that is needed +for the construction of our counit isomorphism `η` -/ +lemma hη : compatibility.τ₀ = + compatibility.τ₁ (eq_to_iso hN₁) (eq_to_iso hΓ₀) + (N₁Γ₀ : Γ ⋙ N₁ ≅ (to_karoubi_equivalence (chain_complex C ℕ)).functor) := +begin + ext K : 3, + simpa only [compatibility.τ₀_hom_app, compatibility.τ₁_hom_app, eq_to_iso.hom, + preadditive.dold_kan.equivalence_counit_iso, N₂Γ₂_to_karoubi_iso_hom, eq_to_hom_map, + eq_to_hom_trans_assoc, eq_to_hom_app] using N₂Γ₂_compatible_with_N₁Γ₀ K, +end + +/-- The counit isomorphism induced by `N₁Γ₀` -/ +@[simps] +def η : Γ ⋙ N ≅ 𝟭 (chain_complex C ℕ) := compatibility.equivalence_counit_iso + (N₁Γ₀ : (Γ : chain_complex C ℕ ⥤ _ ) ⋙ N₁ ≅ (to_karoubi_equivalence _).functor) + +lemma equivalence_counit_iso : + dold_kan.equivalence.counit_iso = (η : Γ ⋙ N ≅ 𝟭 (chain_complex C ℕ)) := +compatibility.equivalence_counit_iso_eq hη + +lemma hε : compatibility.υ (eq_to_iso hN₁) = + (Γ₂N₁ : (to_karoubi_equivalence _).functor ≅ (N₁ : simplicial_object C ⥤ _) ⋙ + preadditive.dold_kan.equivalence.inverse) := +begin + ext X : 4, + erw [nat_trans.comp_app, compatibility_Γ₂N₁_Γ₂N₂_nat_trans], + simp only [compatibility.υ_hom_app, compatibility_Γ₂N₁_Γ₂N₂, + preadditive.dold_kan.equivalence_unit_iso, Γ₂N₂, iso.symm_hom, as_iso_inv, assoc], + erw [← nat_trans.comp_app_assoc, is_iso.hom_inv_id], + dsimp, + simpa only [id_comp, eq_to_hom_app, eq_to_hom_map, eq_to_hom_trans], +end + +/-- The unit isomorphism induced by `Γ₂N₁`. -/ +def ε : 𝟭 (simplicial_object C) ≅ N ⋙ Γ := +compatibility.equivalence_unit_iso (eq_to_iso hΓ₀) Γ₂N₁ + +lemma equivalence_unit_iso : dold_kan.equivalence.unit_iso = + (ε : 𝟭 (simplicial_object C) ≅ N ⋙ Γ) := +compatibility.equivalence_unit_iso_eq hε + +end dold_kan + +end idempotents + +end category_theory diff --git a/src/category_theory/functor/category.lean b/src/category_theory/functor/category.lean index bf73ec5a9f3f1..e3bb5ae20656d 100644 --- a/src/category_theory/functor/category.lean +++ b/src/category_theory/functor/category.lean @@ -60,6 +60,9 @@ lemma congr_app {α β : F ⟶ G} (h : α = β) (X : C) : α.app X = β.app X := @[simp] lemma id_app (F : C ⥤ D) (X : C) : (𝟙 F : F ⟶ F).app X = 𝟙 (F.obj X) := rfl @[simp] lemma comp_app {F G H : C ⥤ D} (α : F ⟶ G) (β : G ⟶ H) (X : C) : (α ≫ β).app X = α.app X ≫ β.app X := rfl +lemma comp_app_assoc {F G H : C ⥤ D} (α : F ⟶ G) (β : G ⟶ H) (X : C) {X' : D} + (f : H.obj X ⟶ X') : + (α ≫ β).app X ≫ f = α.app X ≫ β.app X ≫ f := by rw [comp_app, assoc] lemma app_naturality {F G : C ⥤ (D ⥤ E)} (T : F ⟶ G) (X : C) {Y Z : D} (f : Y ⟶ Z) : ((F.obj X).map f) ≫ ((T.app X).app Z) = ((T.app X).app Y) ≫ ((G.obj X).map f) := From 3ba15165bd6927679be7c22d6091a87337e3cd0c Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 2 Aug 2023 13:20:57 +0000 Subject: [PATCH 1278/1291] feat(analysis/convex/proj_Icc): Extending convex functions (#18797) Constantly extending monotone/antitone functions preserves their convexity. --- src/algebra/order/module.lean | 12 +++ src/algebra/order/monoid/lemmas.lean | 17 ++++ src/algebra/order/smul.lean | 15 ++- src/analysis/convex/proj_Icc.lean | 91 +++++++++++++++++++ src/data/set/intervals/basic.lean | 20 ++++ .../set/intervals/unordered_interval.lean | 25 ++++- src/order/lattice.lean | 40 ++++++++ 7 files changed, 217 insertions(+), 3 deletions(-) create mode 100644 src/analysis/convex/proj_Icc.lean diff --git a/src/algebra/order/module.lean b/src/algebra/order/module.lean index f4b75426245d9..be66690ff1695 100644 --- a/src/algebra/order/module.lean +++ b/src/algebra/order/module.lean @@ -209,6 +209,18 @@ lemma bdd_above.smul_of_nonpos (hc : c ≤ 0) (hs : bdd_above s) : bdd_below (c end ordered_ring +section linear_ordered_ring +variables [linear_ordered_ring k] [linear_ordered_add_comm_group M] [module k M] [ordered_smul k M] + {a : k} + +lemma smul_max_of_nonpos (ha : a ≤ 0) (b₁ b₂ : M) : a • max b₁ b₂ = min (a • b₁) (a • b₂) := +(antitone_smul_left ha : antitone (_ : M → M)).map_max + +lemma smul_min_of_nonpos (ha : a ≤ 0) (b₁ b₂ : M) : a • min b₁ b₂ = max (a • b₁) (a • b₂) := +(antitone_smul_left ha : antitone (_ : M → M)).map_min + +end linear_ordered_ring + section linear_ordered_field variables [linear_ordered_field k] [ordered_add_comm_group M] [module k M] [ordered_smul k M] {s : set M} {c : k} diff --git a/src/algebra/order/monoid/lemmas.lean b/src/algebra/order/monoid/lemmas.lean index 2f2ca1f8b5c17..bbc1e04cab5f1 100644 --- a/src/algebra/order/monoid/lemmas.lean +++ b/src/algebra/order/monoid/lemmas.lean @@ -244,6 +244,23 @@ variables [linear_order α] {a b c d : α} [covariant_class α α (*) (<)] @[to_additive] lemma min_le_max_of_mul_le_mul (h : a * b ≤ c * d) : min a b ≤ max c d := by { simp_rw [min_le_iff, le_max_iff], contrapose! h, exact mul_lt_mul_of_lt_of_lt h.1.1 h.2.2 } +end linear_order + +section linear_order +variables [linear_order α] [covariant_class α α (*) (≤)] [covariant_class α α (swap (*)) (≤)] + {a b c d : α} + +@[to_additive max_add_add_le_max_add_max] lemma max_mul_mul_le_max_mul_max' : + max (a * b) (c * d) ≤ max a c * max b d := +max_le (mul_le_mul' (le_max_left _ _) $ le_max_left _ _) $ + mul_le_mul' (le_max_right _ _) $ le_max_right _ _ + +--TODO: Also missing `min_mul_min_le_min_mul_mul` +@[to_additive min_add_min_le_min_add_add] lemma min_mul_min_le_min_mul_mul' : + min a c * min b d ≤ min (a * b) (c * d) := +le_min (mul_le_mul' (min_le_left _ _) $ min_le_left _ _) $ + mul_le_mul' (min_le_right _ _) $ min_le_right _ _ + end linear_order end has_mul diff --git a/src/algebra/order/smul.lean b/src/algebra/order/smul.lean index b6cd7435b978b..734d7ca1d0213 100644 --- a/src/algebra/order/smul.lean +++ b/src/algebra/order/smul.lean @@ -170,12 +170,23 @@ ordered_smul.mk'' $ λ n hn, begin { cases (int.neg_succ_not_pos _).1 hn } end +section linear_ordered_semiring +variables [linear_ordered_semiring R] [linear_ordered_add_comm_monoid M] [smul_with_zero R M] + [ordered_smul R M] {a : R} + -- TODO: `linear_ordered_field M → ordered_smul ℚ M` -instance linear_ordered_semiring.to_ordered_smul {R : Type*} [linear_ordered_semiring R] : - ordered_smul R R := +instance linear_ordered_semiring.to_ordered_smul : ordered_smul R R := ordered_smul.mk'' $ λ c, strict_mono_mul_left_of_pos +lemma smul_max (ha : 0 ≤ a) (b₁ b₂ : M) : a • max b₁ b₂ = max (a • b₁) (a • b₂) := +(monotone_smul_left ha : monotone (_ : M → M)).map_max + +lemma smul_min (ha : 0 ≤ a) (b₁ b₂ : M) : a • min b₁ b₂ = min (a • b₁) (a • b₂) := +(monotone_smul_left ha : monotone (_ : M → M)).map_min + +end linear_ordered_semiring + section linear_ordered_semifield variables [linear_ordered_semifield 𝕜] [ordered_add_comm_monoid M] [ordered_add_comm_monoid N] [mul_action_with_zero 𝕜 M] [mul_action_with_zero 𝕜 N] diff --git a/src/analysis/convex/proj_Icc.lean b/src/analysis/convex/proj_Icc.lean new file mode 100644 index 0000000000000..3115cdffd85d2 --- /dev/null +++ b/src/analysis/convex/proj_Icc.lean @@ -0,0 +1,91 @@ +/- +Copyright (c) 2023 Yaël Dillies. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies +-/ +import analysis.convex.function +import data.set.intervals.proj_Icc + +/-! +# Convexity of extension from intervals + +This file proves that constantly extending monotone/antitone functions preserves their convexity. + +## TODO + +We could deduplicate the proofs if we had a typeclass stating that `segment 𝕜 x y = [x -[𝕜] y]` as +`𝕜ᵒᵈ` respects it if `𝕜` does, while `𝕜ᵒᵈ` isn't a `linear_ordered_field` if `𝕜` is. +-/ + +open set + +variables {𝕜 β : Type*} [linear_ordered_field 𝕜] [linear_ordered_add_comm_monoid β] [has_smul 𝕜 β] + {s : set 𝕜} {f : 𝕜 → β} {z : 𝕜} + +/-- A convex set extended towards minus infinity is convex. -/ +protected lemma convex.Ici_extend (hf : convex 𝕜 s) : + convex 𝕜 {x | Ici_extend (restrict (Ici z) (∈ s)) x} := +by { rw convex_iff_ord_connected at ⊢ hf, exact hf.restrict.Ici_extend } + +/-- A convex set extended towards infinity is convex. -/ +protected lemma convex.Iic_extend (hf : convex 𝕜 s) : + convex 𝕜 {x | Iic_extend (restrict (Iic z) (∈ s)) x} := +by { rw convex_iff_ord_connected at ⊢ hf, exact hf.restrict.Iic_extend } + +/-- A convex monotone function extended constantly towards minus infinity is convex. -/ +protected lemma convex_on.Ici_extend (hf : convex_on 𝕜 s f) (hf' : monotone_on f s) : + convex_on 𝕜 {x | Ici_extend (restrict (Ici z) (∈ s)) x} (Ici_extend $ restrict (Ici z) f) := +begin + refine ⟨hf.1.Ici_extend, λ x hx y hy a b ha hb hab, _⟩, + dsimp [Ici_extend_apply] at ⊢ hx hy, + refine (hf' (hf.1.ord_connected.uIcc_subset hx hy $ monotone.image_uIcc_subset (λ _ _, max_le_max + le_rfl) $ mem_image_of_mem _ $ convex_uIcc _ _ left_mem_uIcc right_mem_uIcc ha hb hab) + (hf.1 hx hy ha hb hab) _).trans (hf.2 hx hy ha hb hab), + rw [smul_max ha z, smul_max hb z], + refine le_trans _ max_add_add_le_max_add_max, + rw [convex.combo_self hab, smul_eq_mul, smul_eq_mul], +end + +/-- A convex antitone function extended constantly towards infinity is convex. -/ +protected lemma convex_on.Iic_extend (hf : convex_on 𝕜 s f) (hf' : antitone_on f s) : + convex_on 𝕜 {x | Iic_extend (restrict (Iic z) (∈ s)) x} (Iic_extend $ restrict (Iic z) f) := +begin + refine ⟨hf.1.Iic_extend, λ x hx y hy a b ha hb hab, _⟩, + dsimp [Iic_extend_apply] at ⊢ hx hy, + refine (hf' (hf.1 hx hy ha hb hab) (hf.1.ord_connected.uIcc_subset hx hy $ + monotone.image_uIcc_subset (λ _ _, min_le_min le_rfl) $ mem_image_of_mem _ $ + convex_uIcc _ _ left_mem_uIcc right_mem_uIcc ha hb hab) _).trans (hf.2 hx hy ha hb hab), + rw [smul_min ha z, smul_min hb z], + refine min_add_min_le_min_add_add.trans _ , + rw [convex.combo_self hab, smul_eq_mul, smul_eq_mul], +end + +/-- A concave antitone function extended constantly minus towards infinity is concave. -/ +protected lemma concave_on.Ici_extend (hf : concave_on 𝕜 s f) (hf' : antitone_on f s) : + concave_on 𝕜 {x | Ici_extend (restrict (Ici z) (∈ s)) x} (Ici_extend $ restrict (Ici z) f) := +hf.dual.Ici_extend hf'.dual_right + +/-- A concave monotone function extended constantly towards infinity is concave. -/ +protected lemma concave_on.Iic_extend (hf : concave_on 𝕜 s f) (hf' : monotone_on f s) : + concave_on 𝕜 {x | Iic_extend (restrict (Iic z) (∈ s)) x} (Iic_extend $ restrict (Iic z) f) := +hf.dual.Iic_extend hf'.dual_right + +/-- A convex monotone function extended constantly towards minus infinity is convex. -/ +protected lemma convex_on.Ici_extend_of_monotone (hf : convex_on 𝕜 univ f) (hf' : monotone f) : + convex_on 𝕜 univ (Ici_extend $ restrict (Ici z) f) := +hf.Ici_extend $ hf'.monotone_on _ + +/-- A convex antitone function extended constantly towards infinity is convex. -/ +protected lemma convex_on.Iic_extend_of_antitone (hf : convex_on 𝕜 univ f) (hf' : antitone f) : + convex_on 𝕜 univ (Iic_extend $ restrict (Iic z) f) := +hf.Iic_extend $ hf'.antitone_on _ + +/-- A concave antitone function extended constantly minus towards infinity is concave. -/ +protected lemma concave_on.Ici_extend_of_antitone (hf : concave_on 𝕜 univ f) (hf' : antitone f) : + concave_on 𝕜 univ (Ici_extend $ restrict (Ici z) f) := +hf.Ici_extend $ hf'.antitone_on _ + +/-- A concave monotone function extended constantly towards infinity is concave. -/ +protected lemma concave_on.Iic_extend_of_monotone (hf : concave_on 𝕜 univ f) (hf' : monotone f) : + concave_on 𝕜 univ (Iic_extend $ restrict (Iic z) f) := +hf.Iic_extend $ hf'.monotone_on _ diff --git a/src/data/set/intervals/basic.lean b/src/data/set/intervals/basic.lean index 1c4c0e2cc4b6d..250eba029ea68 100644 --- a/src/data/set/intervals/basic.lean +++ b/src/data/set/intervals/basic.lean @@ -1164,6 +1164,26 @@ begin le_of_lt h₂, le_of_lt h₁] }, end +section lattice +variables [lattice β] {f : α → β} + +lemma _root_.monotone_on.image_Icc_subset (hf : monotone_on f (Icc a b)) : + f '' Icc a b ⊆ Icc (f a) (f b) := +image_subset_iff.2 $ λ c hc, + ⟨hf (left_mem_Icc.2 $ hc.1.trans hc.2) hc hc.1, hf hc (right_mem_Icc.2 $ hc.1.trans hc.2) hc.2⟩ + +lemma _root_.antitone_on.image_Icc_subset (hf : antitone_on f (Icc a b)) : + f '' Icc a b ⊆ Icc (f b) (f a) := +image_subset_iff.2 $ λ c hc, + ⟨hf hc (right_mem_Icc.2 $ hc.1.trans hc.2) hc.2, hf (left_mem_Icc.2 $ hc.1.trans hc.2) hc hc.1⟩ + +lemma _root_.monotone.image_Icc_subset (hf : monotone f) : f '' Icc a b ⊆ Icc (f a) (f b) := +(hf.monotone_on _).image_Icc_subset + +lemma _root_.antitone.image_Icc_subset (hf : antitone f) : f '' Icc a b ⊆ Icc (f b) (f a) := +(hf.antitone_on _).image_Icc_subset + +end lattice end linear_order section lattice diff --git a/src/data/set/intervals/unordered_interval.lean b/src/data/set/intervals/unordered_interval.lean index 1e795686afe08..6edbb024832d0 100644 --- a/src/data/set/intervals/unordered_interval.lean +++ b/src/data/set/intervals/unordered_interval.lean @@ -129,7 +129,30 @@ by simpa only [uIcc_comm] using uIcc_injective_right a end distrib_lattice section linear_order -variables [linear_order α] [linear_order β] {f : α → β} {s : set α} {a a₁ a₂ b b₁ b₂ c d x : α} +variables [linear_order α] + +section lattice +variables [lattice β] {f : α → β} {s : set α} {a b : α} + +lemma _root_.monotone_on.image_uIcc_subset (hf : monotone_on f (uIcc a b)) : + f '' uIcc a b ⊆ uIcc (f a) (f b) := +hf.image_Icc_subset.trans $ + by rw [hf.map_sup left_mem_uIcc right_mem_uIcc, hf.map_inf left_mem_uIcc right_mem_uIcc, uIcc] + +lemma _root_.antitone_on.image_uIcc_subset (hf : antitone_on f (uIcc a b)) : + f '' uIcc a b ⊆ uIcc (f a) (f b) := +hf.image_Icc_subset.trans $ + by rw [hf.map_sup left_mem_uIcc right_mem_uIcc, hf.map_inf left_mem_uIcc right_mem_uIcc, uIcc] + +lemma _root_.monotone.image_uIcc_subset (hf : monotone f) : f '' uIcc a b ⊆ uIcc (f a) (f b) := +(hf.monotone_on _).image_uIcc_subset + +lemma _root_.antitone.image_uIcc_subset (hf : antitone f) : f '' uIcc a b ⊆ uIcc (f a) (f b) := +(hf.antitone_on _).image_uIcc_subset + +end lattice + +variables [linear_order β] {f : α → β} {s : set α} {a a₁ a₂ b b₁ b₂ c d x : α} lemma Icc_min_max : Icc (min a b) (max a b) = [a, b] := rfl diff --git a/src/order/lattice.lean b/src/order/lattice.lean index e3877be97db81..14128e9768f95 100644 --- a/src/order/lattice.lean +++ b/src/order/lattice.lean @@ -846,6 +846,7 @@ hf.dual.map_sup _ _ end monotone namespace monotone_on +variables {f : α → β} {s : set α} {x y : α} /-- Pointwise supremum of two monotone functions is a monotone function. -/ protected lemma sup [preorder α] [semilattice_sup β] {f g : α → β} {s : set α} @@ -867,6 +868,25 @@ protected lemma min [preorder α] [linear_order β] {f g : α → β} {s : set (hf : monotone_on f s) (hg : monotone_on g s) : monotone_on (λ x, min (f x) (g x)) s := hf.inf hg +lemma of_map_inf [semilattice_inf α] [semilattice_inf β] + (h : ∀ (x ∈ s) (y ∈ s), f (x ⊓ y) = f x ⊓ f y) : monotone_on f s := +λ x hx y hy hxy, inf_eq_left.1 $ by rw [←h _ hx _ hy, inf_eq_left.2 hxy] + +lemma of_map_sup [semilattice_sup α] [semilattice_sup β] + (h : ∀ (x ∈ s) (y ∈ s), f (x ⊔ y) = f x ⊔ f y) : monotone_on f s := +(@of_map_inf αᵒᵈ βᵒᵈ _ _ _ _ h).dual + +variables [linear_order α] + +lemma map_sup [semilattice_sup β] (hf : monotone_on f s) (hx : x ∈ s) (hy : y ∈ s) : + f (x ⊔ y) = f x ⊔ f y := +by cases le_total x y; have := hf _ _ h; + assumption <|> simp only [h, this, sup_of_le_left, sup_of_le_right] + +lemma map_inf [semilattice_inf β] (hf : monotone_on f s) (hx : x ∈ s) (hy : y ∈ s) : + f (x ⊓ y) = f x ⊓ f y := +hf.dual.map_sup hx hy + end monotone_on namespace antitone @@ -912,6 +932,7 @@ hf.dual_right.map_inf x y end antitone namespace antitone_on +variables {f : α → β} {s : set α} {x y : α} /-- Pointwise supremum of two antitone functions is a antitone function. -/ protected lemma sup [preorder α] [semilattice_sup β] {f g : α → β} {s : set α} @@ -933,6 +954,25 @@ protected lemma min [preorder α] [linear_order β] {f g : α → β} {s : set (hf : antitone_on f s) (hg : antitone_on g s) : antitone_on (λ x, min (f x) (g x)) s := hf.inf hg +lemma of_map_inf [semilattice_inf α] [semilattice_sup β] + (h : ∀ (x ∈ s) (y ∈ s), f (x ⊓ y) = f x ⊔ f y) : antitone_on f s := +λ x hx y hy hxy, sup_eq_left.1 $ by rw [←h _ hx _ hy, inf_eq_left.2 hxy] + +lemma of_map_sup [semilattice_sup α] [semilattice_inf β] + (h : ∀ (x ∈ s) (y ∈ s), f (x ⊔ y) = f x ⊓ f y) : antitone_on f s := +(@of_map_inf αᵒᵈ βᵒᵈ _ _ _ _ h).dual + +variables [linear_order α] + +lemma map_sup [semilattice_inf β] (hf : antitone_on f s) (hx : x ∈ s) (hy : y ∈ s) : + f (x ⊔ y) = f x ⊓ f y := +by cases le_total x y; have := hf _ _ h; assumption <|> + simp only [h, this, sup_of_le_left, sup_of_le_right, inf_of_le_left, inf_of_le_right] + +lemma map_inf [semilattice_sup β] (hf : antitone_on f s) (hx : x ∈ s) (hy : y ∈ s) : + f (x ⊓ y) = f x ⊔ f y := +hf.dual.map_sup hx hy + end antitone_on /-! From 48a058d7e39a80ed56858505719a0b2197900999 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 2 Aug 2023 16:47:14 +0000 Subject: [PATCH 1279/1291] feat(data/sum/interval): The lexicographic sum of two locally finite orders is locally finite (#11352) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This proves `locally_finite_order (α ⊕ₗ β)` where `α` and `β` are locally finite themselves. --- src/data/finset/sum.lean | 2 + src/data/sum/interval.lean | 188 ++++++++++++++++++++++++++++++++++++- 2 files changed, 185 insertions(+), 5 deletions(-) diff --git a/src/data/finset/sum.lean b/src/data/finset/sum.lean index 86d4f7ab8974f..a940205835735 100644 --- a/src/data/finset/sum.lean +++ b/src/data/finset/sum.lean @@ -54,6 +54,8 @@ multiset.mem_disj_sum @[simp] lemma inl_mem_disj_sum : inl a ∈ s.disj_sum t ↔ a ∈ s := inl_mem_disj_sum @[simp] lemma inr_mem_disj_sum : inr b ∈ s.disj_sum t ↔ b ∈ t := inr_mem_disj_sum +@[simp] lemma disj_sum_eq_empty : s.disj_sum t = ∅ ↔ s = ∅ ∧ t = ∅ := by simp [ext_iff] + lemma disj_sum_mono (hs : s₁ ⊆ s₂) (ht : t₁ ⊆ t₂) : s₁.disj_sum t₁ ⊆ s₂.disj_sum t₂ := val_le_iff.1 $ disj_sum_mono (val_le_iff.2 hs) (val_le_iff.2 ht) diff --git a/src/data/sum/interval.lean b/src/data/sum/interval.lean index 5bd433861ad78..e831e4564adfc 100644 --- a/src/data/sum/interval.lean +++ b/src/data/sum/interval.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ +import data.finset.sum import data.sum.order import order.locally_finite @@ -12,11 +13,8 @@ import order.locally_finite > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. > Any changes to this file require a corresponding PR to mathlib4. -This file provides the `locally_finite_order` instance for the disjoint sum of two orders. - -## TODO - -Do the same for the lexicographic sum of orders. +This file provides the `locally_finite_order` instance for the disjoint sum and linear sum of two +orders and calculates the cardinality of their finite intervals. -/ open function sum @@ -99,6 +97,106 @@ lemma sum_lift₂_mono (h₁ : ∀ a b, f₁ a b ⊆ g₁ a b) (h₂ : ∀ a b, | (inr a) (inr b) := map_subset_map.2 (h₂ _ _) end sum_lift₂ + +section sum_lex_lift +variables (f₁ f₁' : α₁ → β₁ → finset γ₁) (f₂ f₂' : α₂ → β₂ → finset γ₂) + (g₁ g₁' : α₁ → β₂ → finset γ₁) (g₂ g₂' : α₁ → β₂ → finset γ₂) + +/-- Lifts maps `α₁ → β₁ → finset γ₁`, `α₂ → β₂ → finset γ₂`, `α₁ → β₂ → finset γ₁`, +`α₂ → β₂ → finset γ₂` to a map `α₁ ⊕ α₂ → β₁ ⊕ β₂ → finset (γ₁ ⊕ γ₂)`. Could be generalized to +alternative monads if we can make sure to keep computability and universe polymorphism. -/ +def sum_lex_lift : Π (a : α₁ ⊕ α₂) (b : β₁ ⊕ β₂), finset (γ₁ ⊕ γ₂) +| (inl a) (inl b) := (f₁ a b).map embedding.inl +| (inl a) (inr b) := (g₁ a b).disj_sum (g₂ a b) +| (inr a) (inl b) := ∅ +| (inr a) (inr b) := (f₂ a b).map ⟨_, inr_injective⟩ + +@[simp] lemma sum_lex_lift_inl_inl (a : α₁) (b : β₁) : + sum_lex_lift f₁ f₂ g₁ g₂ (inl a) (inl b) = (f₁ a b).map embedding.inl := rfl + +@[simp] lemma sum_lex_lift_inl_inr (a : α₁) (b : β₂) : + sum_lex_lift f₁ f₂ g₁ g₂ (inl a) (inr b) = (g₁ a b).disj_sum (g₂ a b) := rfl + +@[simp] lemma sum_lex_lift_inr_inl (a : α₂) (b : β₁) : + sum_lex_lift f₁ f₂ g₁ g₂ (inr a) (inl b) = ∅ := rfl + +@[simp] lemma sum_lex_lift_inr_inr (a : α₂) (b : β₂) : + sum_lex_lift f₁ f₂ g₁ g₂ (inr a) (inr b) = (f₂ a b).map ⟨_, inr_injective⟩ := rfl + +variables {f₁ g₁ f₂ g₂ f₁' g₁' f₂' g₂'} {a : α₁ ⊕ α₂} {b : β₁ ⊕ β₂} {c : γ₁ ⊕ γ₂} + +lemma mem_sum_lex_lift : + c ∈ sum_lex_lift f₁ f₂ g₁ g₂ a b ↔ + (∃ a₁ b₁ c₁, a = inl a₁ ∧ b = inl b₁ ∧ c = inl c₁ ∧ c₁ ∈ f₁ a₁ b₁) ∨ + (∃ a₁ b₂ c₁, a = inl a₁ ∧ b = inr b₂ ∧ c = inl c₁ ∧ c₁ ∈ g₁ a₁ b₂) ∨ + (∃ a₁ b₂ c₂, a = inl a₁ ∧ b = inr b₂ ∧ c = inr c₂ ∧ c₂ ∈ g₂ a₁ b₂) ∨ + ∃ a₂ b₂ c₂, a = inr a₂ ∧ b = inr b₂ ∧ c = inr c₂ ∧ c₂ ∈ f₂ a₂ b₂ := +begin + split, + { cases a; cases b, + { rw [sum_lex_lift, mem_map], + rintro ⟨c, hc, rfl⟩, + exact or.inl ⟨a, b, c, rfl, rfl, rfl, hc⟩ }, + { refine λ h, (mem_disj_sum.1 h).elim _ _, + { rintro ⟨c, hc, rfl⟩, + refine or.inr (or.inl ⟨a, b, c, rfl, rfl, rfl, hc⟩) }, + { rintro ⟨c, hc, rfl⟩, + refine or.inr (or.inr $ or.inl ⟨a, b, c, rfl, rfl, rfl, hc⟩) } }, + { refine λ h, (not_mem_empty _ h).elim }, + { rw [sum_lex_lift, mem_map], + rintro ⟨c, hc, rfl⟩, + exact or.inr (or.inr $ or.inr $ ⟨a, b, c, rfl, rfl, rfl, hc⟩) } }, + { rintro (⟨a, b, c, rfl, rfl, rfl, hc⟩ | ⟨a, b, c, rfl, rfl, rfl, hc⟩ | + ⟨a, b, c, rfl, rfl, rfl, hc⟩ | ⟨a, b, c, rfl, rfl, rfl, hc⟩), + { exact mem_map_of_mem _ hc }, + { exact inl_mem_disj_sum.2 hc }, + { exact inr_mem_disj_sum.2 hc }, + { exact mem_map_of_mem _ hc } } +end + +lemma inl_mem_sum_lex_lift {c₁ : γ₁} : + inl c₁ ∈ sum_lex_lift f₁ f₂ g₁ g₂ a b ↔ + (∃ a₁ b₁, a = inl a₁ ∧ b = inl b₁ ∧ c₁ ∈ f₁ a₁ b₁) ∨ + ∃ a₁ b₂, a = inl a₁ ∧ b = inr b₂ ∧ c₁ ∈ g₁ a₁ b₂ := +by simp [mem_sum_lex_lift] + +lemma inr_mem_sum_lex_lift {c₂ : γ₂} : + inr c₂ ∈ sum_lex_lift f₁ f₂ g₁ g₂ a b ↔ + (∃ a₁ b₂, a = inl a₁ ∧ b = inr b₂ ∧ c₂ ∈ g₂ a₁ b₂) ∨ + ∃ a₂ b₂, a = inr a₂ ∧ b = inr b₂ ∧ c₂ ∈ f₂ a₂ b₂ := +by simp [mem_sum_lex_lift] + +lemma sum_lex_lift_mono (hf₁ : ∀ a b, f₁ a b ⊆ f₁' a b) (hf₂ : ∀ a b, f₂ a b ⊆ f₂' a b) + (hg₁ : ∀ a b, g₁ a b ⊆ g₁' a b) (hg₂ : ∀ a b, g₂ a b ⊆ g₂' a b) (a : α₁ ⊕ α₂) (b : β₁ ⊕ β₂) : + sum_lex_lift f₁ f₂ g₁ g₂ a b ⊆ sum_lex_lift f₁' f₂' g₁' g₂' a b := +begin + cases a; cases b, + exacts [map_subset_map.2 (hf₁ _ _), disj_sum_mono (hg₁ _ _) (hg₂ _ _), subset.rfl, + map_subset_map.2 (hf₂ _ _)], +end + +lemma sum_lex_lift_eq_empty : + (sum_lex_lift f₁ f₂ g₁ g₂ a b) = ∅ ↔ (∀ a₁ b₁, a = inl a₁ → b = inl b₁ → f₁ a₁ b₁ = ∅) ∧ + (∀ a₁ b₂, a = inl a₁ → b = inr b₂ → g₁ a₁ b₂ = ∅ ∧ g₂ a₁ b₂ = ∅) ∧ + ∀ a₂ b₂, a = inr a₂ → b = inr b₂ → f₂ a₂ b₂ = ∅ := +begin + refine ⟨λ h, ⟨_, _, _⟩, λ h, _⟩, + any_goals { rintro a b rfl rfl, exact map_eq_empty.1 h }, + { rintro a b rfl rfl, exact disj_sum_eq_empty.1 h }, + cases a; cases b, + { exact map_eq_empty.2 (h.1 _ _ rfl rfl) }, + { simp [h.2.1 _ _ rfl rfl] }, + { refl }, + { exact map_eq_empty.2 (h.2.2 _ _ rfl rfl) } +end + +lemma sum_lex_lift_nonempty : + (sum_lex_lift f₁ f₂ g₁ g₂ a b).nonempty ↔ (∃ a₁ b₁, a = inl a₁ ∧ b = inl b₁ ∧ (f₁ a₁ b₁).nonempty) + ∨ (∃ a₁ b₂, a = inl a₁ ∧ b = inr b₂ ∧ ((g₁ a₁ b₂).nonempty ∨ (g₂ a₁ b₂).nonempty)) + ∨ ∃ a₂ b₂, a = inr a₂ ∧ b = inr b₂ ∧ (f₂ a₂ b₂).nonempty := +by simp [nonempty_iff_ne_empty, sum_lex_lift_eq_empty, not_and_distrib] + +end sum_lex_lift end finset open finset function @@ -141,4 +239,84 @@ lemma Ioc_inr_inr : Ioc (inr b₁ : α ⊕ β) (inr b₂) = (Ioc b₁ b₂).map lemma Ioo_inr_inr : Ioo (inr b₁ : α ⊕ β) (inr b₂) = (Ioo b₁ b₂).map embedding.inr := rfl end disjoint + +/-! ### Lexicographical sum of orders -/ + +namespace lex +variables [preorder α] [preorder β] [order_top α] [order_bot β] [locally_finite_order α] + [locally_finite_order β] + +/-- Throwaway tactic. -/ +private meta def simp_lex : tactic unit := +`[refine to_lex.surjective.forall₃.2 _, rintro (a | a) (b | b) (c | c); simp only + [sum_lex_lift_inl_inl, sum_lex_lift_inl_inr, sum_lex_lift_inr_inl, sum_lex_lift_inr_inr, + inl_le_inl_iff, inl_le_inr, not_inr_le_inl, inr_le_inr_iff, inl_lt_inl_iff, inl_lt_inr, + not_inr_lt_inl, inr_lt_inr_iff, mem_Icc, mem_Ico, mem_Ioc, mem_Ioo, mem_Ici, mem_Ioi, mem_Iic, + mem_Iio, equiv.coe_to_embedding, to_lex_inj, exists_false, and_false, false_and, map_empty, + not_mem_empty, true_and, inl_mem_disj_sum, inr_mem_disj_sum, and_true, of_lex_to_lex, mem_map, + embedding.coe_fn_mk, exists_prop, exists_eq_right, embedding.inl_apply]] + +instance locally_finite_order : locally_finite_order (α ⊕ₗ β) := +{ finset_Icc := λ a b, + (sum_lex_lift Icc Icc (λ a _, Ici a) (λ _, Iic) (of_lex a) (of_lex b)).map to_lex.to_embedding, + finset_Ico := λ a b, + (sum_lex_lift Ico Ico (λ a _, Ici a) (λ _, Iio) (of_lex a) (of_lex b)).map to_lex.to_embedding, + finset_Ioc := λ a b, + (sum_lex_lift Ioc Ioc (λ a _, Ioi a) (λ _, Iic) (of_lex a) (of_lex b)).map to_lex.to_embedding, + finset_Ioo := λ a b, + (sum_lex_lift Ioo Ioo (λ a _, Ioi a) (λ _, Iio) (of_lex a) (of_lex b)).map to_lex.to_embedding, + finset_mem_Icc := by simp_lex, + finset_mem_Ico := by simp_lex, + finset_mem_Ioc := by simp_lex, + finset_mem_Ioo := by simp_lex } + +variables (a a₁ a₂ : α) (b b₁ b₂ : β) + +lemma Icc_inl_inl : + Icc (inlₗ a₁ : α ⊕ₗ β) (inlₗ a₂) = (Icc a₁ a₂).map (embedding.inl.trans to_lex.to_embedding) := +by { rw ←finset.map_map, refl } + +lemma Ico_inl_inl : + Ico (inlₗ a₁ : α ⊕ₗ β) (inlₗ a₂) = (Ico a₁ a₂).map (embedding.inl.trans to_lex.to_embedding) := +by { rw ←finset.map_map, refl } + +lemma Ioc_inl_inl : + Ioc (inlₗ a₁ : α ⊕ₗ β) (inlₗ a₂) = (Ioc a₁ a₂).map (embedding.inl.trans to_lex.to_embedding) := +by { rw ←finset.map_map, refl } + +lemma Ioo_inl_inl : + Ioo (inlₗ a₁ : α ⊕ₗ β) (inlₗ a₂) = (Ioo a₁ a₂).map (embedding.inl.trans to_lex.to_embedding) := +by { rw ←finset.map_map, refl } + +@[simp] lemma Icc_inl_inr : + Icc (inlₗ a) (inrₗ b) = ((Ici a).disj_sum (Iic b)).map to_lex.to_embedding := rfl +@[simp] lemma Ico_inl_inr : + Ico (inlₗ a) (inrₗ b) = ((Ici a).disj_sum (Iio b)).map to_lex.to_embedding := rfl +@[simp] lemma Ioc_inl_inr : + Ioc (inlₗ a) (inrₗ b) = ((Ioi a).disj_sum (Iic b)).map to_lex.to_embedding := rfl +@[simp] lemma Ioo_inl_inr : + Ioo (inlₗ a) (inrₗ b) = ((Ioi a).disj_sum (Iio b)).map to_lex.to_embedding := rfl + +@[simp] lemma Icc_inr_inl : Icc (inrₗ b) (inlₗ a) = ∅ := rfl +@[simp] lemma Ico_inr_inl : Ico (inrₗ b) (inlₗ a) = ∅ := rfl +@[simp] lemma Ioc_inr_inl : Ioc (inrₗ b) (inlₗ a) = ∅ := rfl +@[simp] lemma Ioo_inr_inl : Ioo (inrₗ b) (inlₗ a) = ∅ := rfl + +lemma Icc_inr_inr : + Icc (inrₗ b₁ : α ⊕ₗ β) (inrₗ b₂) = (Icc b₁ b₂).map (embedding.inr.trans to_lex.to_embedding) := +by { rw ←finset.map_map, refl } + +lemma Ico_inr_inr : + Ico (inrₗ b₁ : α ⊕ₗ β) (inrₗ b₂) = (Ico b₁ b₂).map (embedding.inr.trans to_lex.to_embedding) := +by { rw ←finset.map_map, refl } + +lemma Ioc_inr_inr : + Ioc (inrₗ b₁ : α ⊕ₗ β) (inrₗ b₂) = (Ioc b₁ b₂).map (embedding.inr.trans to_lex.to_embedding) := +by { rw ←finset.map_map, refl } + +lemma Ioo_inr_inr : + Ioo (inrₗ b₁ : α ⊕ₗ β) (inrₗ b₂) = (Ioo b₁ b₂).map (embedding.inr.trans to_lex.to_embedding) := +by { rw ←finset.map_map, refl } + +end lex end sum From 32a7e535287f9c73f2e4d2aef306a39190f0b504 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= <37772949+joelriou@users.noreply.github.com> Date: Sat, 5 Aug 2023 18:07:30 +0000 Subject: [PATCH 1280/1291] feat(algebraic_topology/dold_kan): The Dold-Kan equivalence for abelian categories (#17926) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Joël Riou --- .../dold_kan/compatibility.lean | 4 +- .../dold_kan/decomposition.lean | 2 + .../dold_kan/degeneracies.lean | 2 + .../dold_kan/equivalence.lean | 175 ++++++++++++++++++ .../dold_kan/equivalence_additive.lean | 2 + .../dold_kan/equivalence_pseudoabelian.lean | 2 + src/algebraic_topology/dold_kan/faces.lean | 2 + .../dold_kan/functor_gamma.lean | 2 + .../dold_kan/functor_n.lean | 2 +- .../dold_kan/gamma_comp_n.lean | 2 + .../dold_kan/n_comp_gamma.lean | 2 + .../dold_kan/n_reflects_iso.lean | 2 + .../dold_kan/normalized.lean | 2 + .../dold_kan/notations.lean | 2 + src/algebraic_topology/dold_kan/p_infty.lean | 3 +- .../dold_kan/projections.lean | 2 + .../dold_kan/split_simplicial_object.lean | 3 + 17 files changed, 208 insertions(+), 3 deletions(-) create mode 100644 src/algebraic_topology/dold_kan/equivalence.lean diff --git a/src/algebraic_topology/dold_kan/compatibility.lean b/src/algebraic_topology/dold_kan/compatibility.lean index 3567efda0ccc9..516e38ecaf374 100644 --- a/src/algebraic_topology/dold_kan/compatibility.lean +++ b/src/algebraic_topology/dold_kan/compatibility.lean @@ -36,7 +36,9 @@ inverse of `eB`: but whose inverse functor is `G`. When extra assumptions are given, we shall also provide simplification lemmas for the -unit and counit isomorphisms of `equivalence`. (TODO) +unit and counit isomorphisms of `equivalence`. + +(See `equivalence.lean` for the general strategy of proof of the Dold-Kan equivalence.) -/ diff --git a/src/algebraic_topology/dold_kan/decomposition.lean b/src/algebraic_topology/dold_kan/decomposition.lean index 1d89283129af6..f1cbbb85fd8b7 100644 --- a/src/algebraic_topology/dold_kan/decomposition.lean +++ b/src/algebraic_topology/dold_kan/decomposition.lean @@ -29,6 +29,8 @@ role in the proof that the functor `N₁ : simplicial_object C ⥤ karoubi (chain_complex C ℕ))` reflects isomorphisms. +(See `equivalence.lean` for the general strategy of proof of the Dold-Kan equivalence.) + -/ open category_theory category_theory.category category_theory.preadditive opposite diff --git a/src/algebraic_topology/dold_kan/degeneracies.lean b/src/algebraic_topology/dold_kan/degeneracies.lean index c3b3de9892ba2..ab0584b061494 100644 --- a/src/algebraic_topology/dold_kan/degeneracies.lean +++ b/src/algebraic_topology/dold_kan/degeneracies.lean @@ -25,6 +25,8 @@ if `X : simplicial_object C` with `C` a preadditive category, `X.map θ.op ≫ P_infty.f n = 0`. It follows from the more precise statement vanishing statement `σ_comp_P_eq_zero` for the `P q`. +(See `equivalence.lean` for the general strategy of proof of the Dold-Kan equivalence.) + -/ open category_theory category_theory.category category_theory.limits diff --git a/src/algebraic_topology/dold_kan/equivalence.lean b/src/algebraic_topology/dold_kan/equivalence.lean new file mode 100644 index 0000000000000..dc025a7938b2b --- /dev/null +++ b/src/algebraic_topology/dold_kan/equivalence.lean @@ -0,0 +1,175 @@ +/- +Copyright (c) 2022 Joël Riou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joël Riou +-/ + +import algebraic_topology.dold_kan.equivalence_pseudoabelian +import algebraic_topology.dold_kan.normalized + +/-! + +# The Dold-Kan correspondence + +The Dold-Kan correspondence states that for any abelian category `A`, there is +an equivalence between the category of simplicial objects in `A` and the +category of chain complexes in `A` (with degrees indexed by `ℕ` and the +homological convention that the degree is decreased by the differentials). + +In this file, we finish the construction of this equivalence by providing +`category_theory.abelian.dold_kan.equivalence` which is of type +`simplicial_object A ≌ chain_complex A ℕ` for any abelian category `A`. +The functor `simplicial_object A ⥤ chain_complex A ℕ` of this equivalence is +definitionally equal to `normalized_Moore_complex A`. + +## Overall strategy of the proof of the correspondence + +Before starting the implementation of the proof in Lean, the author noticed +that the Dold-Kan equivalence not only applies to abelian categories, but +should also hold generally for any pseudoabelian category `C` +(i.e. a category with instances `[preadditive C]` +`[has_finite_coproducts C]` and `[is_idempotent_complete C]`): this is +`category_theory.idempotents.dold_kan.equivalence`. + +When the alternating face map complex `K[X]` of a simplicial object `X` in an +abelian is studied, it is shown that it decomposes as a direct sum of the +normalized subcomplex and of the degenerate subcomplex. The crucial observation +is that in this decomposition, the projection on the normalized subcomplex can +be defined in each degree using simplicial operators. Then, the definition +of this projection `P_infty : K[X] ⟶ K[X]` can be carried out for any +`X : simplicial_object C` when `C` is a preadditive category. + +The construction of the endomorphism `P_infty` is done in the files +`homotopies.lean`, `faces.lean`, `projections.lean` and `p_infty.lean`. +Eventually, as we would also like to show that the inclusion of the normalized +Moore complex is a homotopy equivalence (cf. file `homotopy_equivalence.lean`), +this projection `P_infty` needs to be homotopic to the identity. In our +construction, we get this for free because `P_infty` is obtained by altering +the identity endomorphism by null homotopic maps. More details about this +aspect of the proof are in the file `homotopies.lean`. + +When the alternating face map complex `K[X]` is equipped with the idempotent +endomorphism `P_infty`, it becomes an object in `karoubi (chain_complex C ℕ)` +which is the idempotent completion of the category `chain_complex C ℕ`. In `functor_n.lean`, +we obtain this functor `N₁ : simplicial_object C ⥤ karoubi (chain_complex C ℕ)`, +which is formally extended as +`N₂ : karoubi (simplicial_object C) ⥤ karoubi (chain_complex C ℕ)`. (Here, some functors +have an index which is the number of occurrences of `karoubi` at the source or the +target.) + +In `functor_gamma.lean`, assuming that the category `C` is additive, +we define the functor in the other direction +`Γ₂ : karoubi (chain_complex C ℕ) ⥤ karoubi (simplicial_object C)` as the formal +extension of a functor `Γ₀ : chain_complex C ℕ ⥤ simplicial_object C` which is +defined similarly as in *Simplicial Homotopy Theory* by Goerss-Jardine. +In `degeneracies.lean`, we show that `P_infty` vanishes on the image of degeneracy +operators, which is one of the key properties that makes it possible to contruct +the isomorphism `N₂Γ₂ : Γ₂ ⋙ N₂ ≅ 𝟭 (karoubi (chain_complex C ℕ))`. + +The rest of the proof follows the strategy in the [original paper by Dold][dold1958]. We show +that the functor `N₂` reflects isomorphisms in `n_reflects_iso.lean`: this relies on a +decomposition of the identity of `X _[n]` using `P_infty.f n` and degeneracies obtained in +`decomposition.lean`. Then, in `n_comp_gamma.lean`, we construct a natural transformation +`Γ₂N₂.trans : N₂ ⋙ Γ₂ ⟶ 𝟭 (karoubi (simplicial_object C))`. It is shown that it is an +isomorphism using the fact that `N₂` reflects isomorphisms, and because we can show +that the composition `N₂ ⟶ N₂ ⋙ Γ₂ ⋙ N₂ ⟶ N₂` is the identity (see `identity_N₂`). The fact +that `N₂` is defined as a formal direct factor makes the proof easier because we only +have to compare endomorphisms of an alternating face map complex `K[X]` and we do not +have to worry with inclusions of kernel subobjects. + +In `equivalence_additive.lean`, we obtain +the equivalence `equivalence : karoubi (simplicial_object C) ≌ karoubi (chain_complex C ℕ)`. +It is in the namespace `category_theory.preadditive.dold_kan`. The functors in this +equivalence are named `N` and `Γ`: by definition, they are `N₂` and `Γ₂`. + +In `equivalence_pseudoabelian.lean`, assuming `C` is idempotent complete, +we obtain `equivalence : simplicial_object C ≌ chain_complex C ℕ` +in the namespace `category_theory.idempotents.dold_kan`. This could be roughly +obtained by composing the previous equivalence with the equivalences +`simplicial_object C ≌ karoubi (simplicial_object C)` and +`karoubi (chain_complex C ℕ) ≌ chain_complex C ℕ`. Instead, we polish this construction +in `compatibility.lean` by ensuring good definitional properties of the equivalence (e.g. +the inverse functor is definitionallly equal to +`Γ₀' : chain_complex C ℕ ⥤ simplicial_object C`) and +showing compatibilities for the unit and counit isomorphisms. + +In this file `equivalence.lean`, assuming the category `A` is abelian, we obtain +`equivalence : simplicial_object A ≌ chain_complex A ℕ` in the namespace +`category_theory.abelian.dold_kan`. This is obtained by replacing the functor +`category_theory.idempotents.dold_kan.N` of the equivalence in the pseudoabelian case +with the isomorphic functor `normalized_Moore_complex A` thanks to the isomorphism +obtained in `normalized.lean`. + +TODO: Show functoriality properties of the three equivalences above. More precisely, +for example in the case of abelian categories `A` and `B`, if `F : A ⥤ B` is an +additive functor, we can show that the functors `N` for `A` and `B` are compatible +with the functors `simplicial_object A ⥤ simplicial_object B` and +`chain_complex A ℕ ⥤ chain_complex B ℕ` induced by `F`. (Note that this does not +require that `F` is an exact functor!) + +TODO: Introduce the degenerate subcomplex `D[X]` which is generated by +degenerate simplices, show that the projector `P_infty` corresponds to +a decomposition `K[X] ≅ N[X] ⊞ D[X]`. + +TODO: dualise all of this as `cosimplicial_object A ⥤ cochain_complex A ℕ`. (It is unclear +what is the best way to do this. The exact design may be decided when it is needed.) + +## References +* [Albrecht Dold, Homology of Symmetric Products and Other Functors of Complexes][dold1958] +* [Paul G. Goerss, John F. Jardine, Simplicial Homotopy Theory][goerss-jardine-2009] + +-/ + +noncomputable theory + +open category_theory +open category_theory.category +open category_theory.idempotents + +variables {A : Type*} [category A] [abelian A] + +namespace category_theory + +namespace abelian + +namespace dold_kan + +open algebraic_topology.dold_kan + +/-- The functor `N` for the equivalence is `normalized_Moore_complex A` -/ +def N : simplicial_object A ⥤ chain_complex A ℕ := algebraic_topology.normalized_Moore_complex A + +/-- The functor `Γ` for the equivalence is the same as in the pseudoabelian case. -/ +def Γ : chain_complex A ℕ ⥤ simplicial_object A := idempotents.dold_kan.Γ + +/-- The comparison isomorphism between `normalized_Moore_complex A` and +the functor `idempotents.dold_kan.N` from the pseudoabelian case -/ +@[simps] +def comparison_N : (N : simplicial_object A ⥤ _) ≅ idempotents.dold_kan.N := +calc N ≅ N ⋙ 𝟭 _ : functor.left_unitor N +... ≅ N ⋙ ((to_karoubi_equivalence _).functor ⋙ (to_karoubi_equivalence _).inverse) : + iso_whisker_left _ (to_karoubi_equivalence _).unit_iso +... ≅ (N ⋙ (to_karoubi_equivalence _).functor) ⋙ (to_karoubi_equivalence _).inverse : + iso.refl _ +... ≅ N₁ ⋙ (to_karoubi_equivalence _).inverse : iso_whisker_right + (N₁_iso_normalized_Moore_complex_comp_to_karoubi A).symm _ +... ≅ idempotents.dold_kan.N : by refl + +/-- The Dold-Kan equivalence for abelian categories -/ +@[simps functor] +def equivalence : simplicial_object A ≌ chain_complex A ℕ := +begin + let F : simplicial_object A ⥤ _ := idempotents.dold_kan.N, + let hF : is_equivalence F := is_equivalence.of_equivalence idempotents.dold_kan.equivalence, + letI : is_equivalence (N : simplicial_object A ⥤ _ ) := + is_equivalence.of_iso comparison_N.symm hF, + exact N.as_equivalence, +end + +lemma equivalence_inverse : (equivalence : simplicial_object A ≌ _).inverse = Γ := rfl + +end dold_kan + +end abelian + +end category_theory diff --git a/src/algebraic_topology/dold_kan/equivalence_additive.lean b/src/algebraic_topology/dold_kan/equivalence_additive.lean index 886aefc7db66b..742196ac2a72b 100644 --- a/src/algebraic_topology/dold_kan/equivalence_additive.lean +++ b/src/algebraic_topology/dold_kan/equivalence_additive.lean @@ -14,6 +14,8 @@ import algebraic_topology.dold_kan.n_comp_gamma This file defines `preadditive.dold_kan.equivalence` which is the equivalence of categories `karoubi (simplicial_object C) ≌ karoubi (chain_complex C ℕ)`. +(See `equivalence.lean` for the general strategy of proof of the Dold-Kan equivalence.) + -/ noncomputable theory diff --git a/src/algebraic_topology/dold_kan/equivalence_pseudoabelian.lean b/src/algebraic_topology/dold_kan/equivalence_pseudoabelian.lean index dd766e8a5e41f..4d6fff16434a8 100644 --- a/src/algebraic_topology/dold_kan/equivalence_pseudoabelian.lean +++ b/src/algebraic_topology/dold_kan/equivalence_pseudoabelian.lean @@ -29,6 +29,8 @@ the composition of `N₁ : simplicial_object C ⥤ karoubi (chain_complex C ℕ) `idempotents.dold_kan.Γ` of the equivalence is by definition the functor `Γ₀` introduced in `functor_gamma.lean`. +(See `equivalence.lean` for the general strategy of proof of the Dold-Kan equivalence.) + -/ noncomputable theory diff --git a/src/algebraic_topology/dold_kan/faces.lean b/src/algebraic_topology/dold_kan/faces.lean index ae2b2f6e8fff0..d9e7af0e63960 100644 --- a/src/algebraic_topology/dold_kan/faces.lean +++ b/src/algebraic_topology/dold_kan/faces.lean @@ -25,6 +25,8 @@ The main lemma in this file is `higher_faces_vanish.induction`. It is based on two technical lemmas `higher_faces_vanish.comp_Hσ_eq` and `higher_faces_vanish.comp_Hσ_eq_zero`. +(See `equivalence.lean` for the general strategy of proof of the Dold-Kan equivalence.) + -/ open nat diff --git a/src/algebraic_topology/dold_kan/functor_gamma.lean b/src/algebraic_topology/dold_kan/functor_gamma.lean index 9a5675b665216..bd138e8ba7d11 100644 --- a/src/algebraic_topology/dold_kan/functor_gamma.lean +++ b/src/algebraic_topology/dold_kan/functor_gamma.lean @@ -27,6 +27,8 @@ By construction, `Γ₀.obj K` is a split simplicial object whose splitting is ` We also construct `Γ₂ : karoubi (chain_complex C ℕ) ⥤ karoubi (simplicial_object C)` which shall be an equivalence for any additive category `C`. +(See `equivalence.lean` for the general strategy of proof of the Dold-Kan equivalence.) + -/ noncomputable theory diff --git a/src/algebraic_topology/dold_kan/functor_n.lean b/src/algebraic_topology/dold_kan/functor_n.lean index 17a1db8bb19eb..74f4d58a98df4 100644 --- a/src/algebraic_topology/dold_kan/functor_n.lean +++ b/src/algebraic_topology/dold_kan/functor_n.lean @@ -32,7 +32,7 @@ defined in `equivalence_pseudoabelian.lean`. When the category `C` is abelian, a relation between `N₁` and the normalized Moore complex functor shall be obtained in `normalized.lean`. -(See `equivalence.lean` for the general strategy of proof.) +(See `equivalence.lean` for the general strategy of proof of the Dold-Kan equivalence.) -/ diff --git a/src/algebraic_topology/dold_kan/gamma_comp_n.lean b/src/algebraic_topology/dold_kan/gamma_comp_n.lean index b3d66c5c2cfcc..86e9a3e0230c0 100644 --- a/src/algebraic_topology/dold_kan/gamma_comp_n.lean +++ b/src/algebraic_topology/dold_kan/gamma_comp_n.lean @@ -16,6 +16,8 @@ The purpose of this file is to construct natural isomorphisms `N₁Γ₀ : Γ₀ ⋙ N₁ ≅ to_karoubi (chain_complex C ℕ)` and `N₂Γ₂ : Γ₂ ⋙ N₂ ≅ 𝟭 (karoubi (chain_complex C ℕ))`. +(See `equivalence.lean` for the general strategy of proof of the Dold-Kan equivalence.) + -/ noncomputable theory diff --git a/src/algebraic_topology/dold_kan/n_comp_gamma.lean b/src/algebraic_topology/dold_kan/n_comp_gamma.lean index 52bdb21046dc0..ac2c1537c3539 100644 --- a/src/algebraic_topology/dold_kan/n_comp_gamma.lean +++ b/src/algebraic_topology/dold_kan/n_comp_gamma.lean @@ -21,6 +21,8 @@ that it becomes an isomorphism after the application of the functor `N₂ : karoubi (simplicial_object C) ⥤ karoubi (chain_complex C ℕ)` which reflects isomorphisms. +(See `equivalence.lean` for the general strategy of proof of the Dold-Kan equivalence.) + -/ noncomputable theory diff --git a/src/algebraic_topology/dold_kan/n_reflects_iso.lean b/src/algebraic_topology/dold_kan/n_reflects_iso.lean index d9cd2e612a1e0..10541d6b39f98 100644 --- a/src/algebraic_topology/dold_kan/n_reflects_iso.lean +++ b/src/algebraic_topology/dold_kan/n_reflects_iso.lean @@ -21,6 +21,8 @@ In this file, it is shown that the functors `N₂ : karoubi (simplicial_object C) ⥤ karoubi (chain_complex C ℕ))` reflect isomorphisms for any preadditive category `C`. +(See `equivalence.lean` for the general strategy of proof of the Dold-Kan equivalence.) + -/ open category_theory diff --git a/src/algebraic_topology/dold_kan/normalized.lean b/src/algebraic_topology/dold_kan/normalized.lean index c182615e41b7c..d01fa53ec4cd0 100644 --- a/src/algebraic_topology/dold_kan/normalized.lean +++ b/src/algebraic_topology/dold_kan/normalized.lean @@ -27,6 +27,8 @@ the Dold-Kan equivalence `category_theory.abelian.dold_kan.equivalence : simplicial_object A ≌ chain_complex A ℕ` with a functor (definitionally) equal to `normalized_Moore_complex A`. +(See `equivalence.lean` for the general strategy of proof of the Dold-Kan equivalence.) + -/ open category_theory category_theory.category category_theory.limits diff --git a/src/algebraic_topology/dold_kan/notations.lean b/src/algebraic_topology/dold_kan/notations.lean index 67fda8f17e311..3276ab335e94a 100644 --- a/src/algebraic_topology/dold_kan/notations.lean +++ b/src/algebraic_topology/dold_kan/notations.lean @@ -17,6 +17,8 @@ This file defines the notation `K[X] : chain_complex C ℕ` for the alternating map complex of `(X : simplicial_object C)` where `C` is a preadditive category, as well as `N[X]` for the normalized subcomplex in the case `C` is an abelian category. +(See `equivalence.lean` for the general strategy of proof of the Dold-Kan equivalence.) + -/ localized "notation (name := alternating_face_map_complex) `K[`X`]` := diff --git a/src/algebraic_topology/dold_kan/p_infty.lean b/src/algebraic_topology/dold_kan/p_infty.lean index b95df628099e3..7b1c539dca8bf 100644 --- a/src/algebraic_topology/dold_kan/p_infty.lean +++ b/src/algebraic_topology/dold_kan/p_infty.lean @@ -22,7 +22,8 @@ to the limit the projections `P q` defined in `projections.lean`. This projection is a critical tool in this formalisation of the Dold-Kan correspondence, because in the case of abelian categories, `P_infty` corresponds to the projection on the normalized Moore subcomplex, with kernel the degenerate subcomplex. -(See `equivalence.lean` for the general strategy of proof.) + +(See `equivalence.lean` for the general strategy of proof of the Dold-Kan equivalence.) -/ diff --git a/src/algebraic_topology/dold_kan/projections.lean b/src/algebraic_topology/dold_kan/projections.lean index 2439974a9359a..4b9ea21fdc872 100644 --- a/src/algebraic_topology/dold_kan/projections.lean +++ b/src/algebraic_topology/dold_kan/projections.lean @@ -30,6 +30,8 @@ By passing to the limit, these endomorphisms `P q` shall be used in `p_infty.lea in order to define `P_infty : K[X] ⟶ K[X]`, see `equivalence.lean` for the general strategy of proof of the Dold-Kan equivalence. +(See `equivalence.lean` for the general strategy of proof of the Dold-Kan equivalence.) + -/ open category_theory category_theory.category category_theory.limits diff --git a/src/algebraic_topology/dold_kan/split_simplicial_object.lean b/src/algebraic_topology/dold_kan/split_simplicial_object.lean index f3b2c7fec482f..997675cde1821 100644 --- a/src/algebraic_topology/dold_kan/split_simplicial_object.lean +++ b/src/algebraic_topology/dold_kan/split_simplicial_object.lean @@ -18,6 +18,9 @@ import algebraic_topology.dold_kan.functor_n In this file we define a functor `nondeg_complex : simplicial_object.split C ⥤ chain_complex C ℕ` when `C` is a preadditive category with finite coproducts, and get an isomorphism `to_karoubi_nondeg_complex_iso_N₁ : nondeg_complex ⋙ to_karoubi _ ≅ forget C ⋙ dold_kan.N₁`. + +(See `equivalence.lean` for the general strategy of proof of the Dold-Kan equivalence.) + -/ noncomputable theory From ffde2d8a6e689149e44fd95fa862c23a57f8c780 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 28 Aug 2023 11:58:43 +0000 Subject: [PATCH 1281/1291] chore(order/liminf_limsup): Generalise and move lemmas (#18628) Generalise lemmas from semilattices to codirected orders. Move topology-less lemmas from `topology.algebra.order.liminf_limsup` to `order.liminf_limsup`. Also turn arguments to `bdd_above_insert` and friends implicit. --- src/data/set/finite.lean | 17 ++-- src/order/bounds/basic.lean | 77 +++++++++++++------ src/order/directed.lean | 40 ---------- src/order/liminf_limsup.lean | 46 +++++++++-- src/topology/algebra/order/liminf_limsup.lean | 35 --------- 5 files changed, 99 insertions(+), 116 deletions(-) diff --git a/src/data/set/finite.lean b/src/data/set/finite.lean index 67eb58978f762..f86cca6bfa443 100644 --- a/src/data/set/finite.lean +++ b/src/data/set/finite.lean @@ -1280,7 +1280,7 @@ end section -variables [semilattice_sup α] [nonempty α] {s : set α} +variables [preorder α] [is_directed α (≤)] [nonempty α] {s : set α} /--A finite set is bounded above.-/ protected lemma finite.bdd_above (hs : s.finite) : bdd_above s := @@ -1288,7 +1288,7 @@ finite.induction_on hs bdd_above_empty $ λ a s _ _ h, h.insert a /--A finite union of sets which are all bounded above is still bounded above.-/ lemma finite.bdd_above_bUnion {I : set β} {S : β → set α} (H : I.finite) : - (bdd_above (⋃i∈I, S i)) ↔ (∀i ∈ I, bdd_above (S i)) := + bdd_above (⋃ i ∈ I, S i) ↔ ∀ i ∈ I, bdd_above (S i) := finite.induction_on H (by simp only [bUnion_empty, bdd_above_empty, ball_empty_iff]) (λ a s ha _ hs, by simp only [bUnion_insert, ball_insert_iff, bdd_above_union, hs]) @@ -1299,22 +1299,17 @@ end section -variables [semilattice_inf α] [nonempty α] {s : set α} +variables [preorder α] [is_directed α (≥)] [nonempty α] {s : set α} /--A finite set is bounded below.-/ -protected lemma finite.bdd_below (hs : s.finite) : bdd_below s := @finite.bdd_above αᵒᵈ _ _ _ hs +protected lemma finite.bdd_below (hs : s.finite) : 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 : I.finite) : bdd_below (⋃ i ∈ I, S i) ↔ ∀ i ∈ I, bdd_below (S i) := -@finite.bdd_above_bUnion αᵒᵈ _ _ _ _ _ H +@finite.bdd_above_bUnion αᵒᵈ _ _ _ _ _ _ H -lemma infinite_of_not_bdd_below : ¬ bdd_below s → s.infinite := -begin - contrapose!, - rw not_infinite, - apply finite.bdd_below, -end +lemma infinite_of_not_bdd_below : ¬ bdd_below s → s.infinite := mt finite.bdd_below end diff --git a/src/order/bounds/basic.lean b/src/order/bounds/basic.lean index 6cfd3c71db1d3..6add0bb14d555 100644 --- a/src/order/bounds/basic.lean +++ b/src/order/bounds/basic.lean @@ -5,6 +5,7 @@ Authors: Johannes Hölzl, Yury Kudryashov -/ import data.set.intervals.basic import data.set.n_ary +import order.directed /-! @@ -283,31 +284,30 @@ h.mono $ inter_subset_left s t lemma bdd_below.inter_of_right (h : bdd_below t) : bdd_below (s ∩ t) := h.mono $ inter_subset_right s t -/-- If `s` and `t` are bounded above sets in a `semilattice_sup`, then so is `s ∪ t`. -/ -lemma bdd_above.union [semilattice_sup γ] {s t : set γ} : +/-- In a directed order, the union of bounded above sets is bounded above. -/ +lemma bdd_above.union [is_directed α (≤)] {s t : set α} : bdd_above s → bdd_above t → bdd_above (s ∪ t) := begin - rintros ⟨bs, hs⟩ ⟨bt, ht⟩, - use bs ⊔ bt, - rw upper_bounds_union, - exact ⟨upper_bounds_mono_mem le_sup_left hs, - upper_bounds_mono_mem le_sup_right ht⟩ + rintro ⟨a, ha⟩ ⟨b, hb⟩, + obtain ⟨c, hca, hcb⟩ := exists_ge_ge a b, + rw [bdd_above, upper_bounds_union], + exact ⟨c, upper_bounds_mono_mem hca ha, upper_bounds_mono_mem hcb hb⟩, end -/-- The union of two sets is bounded above if and only if each of the sets is. -/ -lemma bdd_above_union [semilattice_sup γ] {s t : set γ} : +/-- In a directed order, the union of two sets is bounded above if and only if both sets are. -/ +lemma bdd_above_union [is_directed α (≤)] {s t : set α} : bdd_above (s ∪ t) ↔ bdd_above s ∧ bdd_above t := -⟨λ h, ⟨h.mono $ subset_union_left s t, h.mono $ subset_union_right s t⟩, - λ h, h.1.union h.2⟩ +⟨λ h, ⟨h.mono $ subset_union_left _ _, h.mono $ subset_union_right _ _⟩, λ h, h.1.union h.2⟩ -lemma bdd_below.union [semilattice_inf γ] {s t : set γ} : +/-- In a codirected order, the union of bounded below sets is bounded below. -/ +lemma bdd_below.union [is_directed α (≥)] {s t : set α} : bdd_below s → bdd_below t → bdd_below (s ∪ t) := -@bdd_above.union γᵒᵈ _ s t +@bdd_above.union αᵒᵈ _ _ _ _ -/--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 γ} : +/-- In a codirected order, the union of two sets is bounded below if and only if both sets are. -/ +lemma bdd_below_union [is_directed α (≥)] {s t : set α} : bdd_below (s ∪ t) ↔ bdd_below s ∧ bdd_below t := -@bdd_above_union γᵒᵈ _ s t +@bdd_above_union αᵒᵈ _ _ _ _ /-- 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`. -/ @@ -642,22 +642,22 @@ lemma nonempty_of_not_bdd_below [ha : nonempty α] (h : ¬bdd_below s) : s.nonem -/ /-- Adding a point to a set preserves its boundedness above. -/ -@[simp] lemma bdd_above_insert [semilattice_sup γ] (a : γ) {s : set γ} : +@[simp] lemma bdd_above_insert [is_directed α (≤)] {s : set α} {a : α} : bdd_above (insert a s) ↔ bdd_above s := by simp only [insert_eq, bdd_above_union, bdd_above_singleton, true_and] -lemma bdd_above.insert [semilattice_sup γ] (a : γ) {s : set γ} (hs : bdd_above s) : - bdd_above (insert a s) := -(bdd_above_insert a).2 hs +protected lemma bdd_above.insert [is_directed α (≤)] {s : set α} (a : α) : + bdd_above s → bdd_above (insert a s) := +bdd_above_insert.2 /--Adding a point to a set preserves its boundedness below.-/ -@[simp] lemma bdd_below_insert [semilattice_inf γ] (a : γ) {s : set γ} : +@[simp] lemma bdd_below_insert [is_directed α (≥)] {s : set α} {a : α} : bdd_below (insert a s) ↔ bdd_below s := by simp only [insert_eq, bdd_below_union, bdd_below_singleton, true_and] -lemma bdd_below.insert [semilattice_inf γ] (a : γ) {s : set γ} (hs : bdd_below s) : - bdd_below (insert a s) := -(bdd_below_insert a).2 hs +lemma bdd_below.insert [is_directed α (≥)] {s : set α} (a : α) : + bdd_below s → bdd_below (insert a s) := +bdd_below_insert.2 lemma is_lub.insert [semilattice_sup γ] (a) {b} {s : set γ} (hs : is_lub s b) : is_lub (insert a s) (a ⊔ b) := @@ -1191,3 +1191,32 @@ 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 αᵒᵈ βᵒᵈ _ _ _ _ + +section scott_continuous +variables [preorder α] [preorder β] {f : α → β} {a : α} + +/-- A function between preorders is said to be Scott continuous if it preserves `is_lub` on directed +sets. It can be shown that a function is Scott continuous if and only if it is continuous wrt the +Scott topology. + +The dual notion + +```lean +∀ ⦃d : set α⦄, d.nonempty → directed_on (≥) d → ∀ ⦃a⦄, is_glb d a → is_glb (f '' d) (f a) +``` + +does not appear to play a significant role in the literature, so is omitted here. +-/ +def scott_continuous (f : α → β) : Prop := +∀ ⦃d : set α⦄, d.nonempty → directed_on (≤) d → ∀ ⦃a⦄, is_lub d a → is_lub (f '' d) (f a) + +protected lemma scott_continuous.monotone (h : scott_continuous f) : monotone f := +begin + refine λ a b hab, (h (insert_nonempty _ _) (directed_on_pair le_refl hab) _).1 + (mem_image_of_mem _ $ mem_insert _ _), + rw [is_lub, upper_bounds_insert, upper_bounds_singleton, + inter_eq_self_of_subset_right (Ici_subset_Ici.2 hab)], + exact is_least_Ici, +end + +end scott_continuous diff --git a/src/order/directed.lean b/src/order/directed.lean index 83c44b4bf2b62..0456ab1e72b13 100644 --- a/src/order/directed.lean +++ b/src/order/directed.lean @@ -6,7 +6,6 @@ Authors: Johannes Hölzl import data.set.image import order.lattice import order.max -import order.bounds.basic /-! # Directed indexed families and sets @@ -259,42 +258,3 @@ instance order_top.to_is_directed_le [has_le α] [order_top α] : is_directed α @[priority 100] -- see Note [lower instance priority] instance order_bot.to_is_directed_ge [has_le α] [order_bot α] : is_directed α (≥) := ⟨λ a b, ⟨⊥, bot_le, bot_le⟩⟩ - -section scott_continuous - -variables [preorder α] {a : α} - -/-- -A function between preorders is said to be Scott continuous if it preserves `is_lub` on directed -sets. It can be shown that a function is Scott continuous if and only if it is continuous wrt the -Scott topology. - -The dual notion - -```lean -∀ ⦃d : set α⦄, d.nonempty → directed_on (≥) d → ∀ ⦃a⦄, is_glb d a → is_glb (f '' d) (f a) -``` - -does not appear to play a significant role in the literature, so is omitted here. --/ -def scott_continuous [preorder β] (f : α → β) : Prop := -∀ ⦃d : set α⦄, d.nonempty → directed_on (≤) d → ∀ ⦃a⦄, is_lub d a → is_lub (f '' d) (f a) - -protected lemma scott_continuous.monotone [preorder β] {f : α → β} - (h : scott_continuous f) : - monotone f := -begin - intros a b hab, - have e1 : is_lub (f '' {a, b}) (f b), - { apply h, - { exact set.insert_nonempty _ _ }, - { exact directed_on_pair le_refl hab }, - { rw [is_lub, upper_bounds_insert, upper_bounds_singleton, - set.inter_eq_self_of_subset_right (set.Ici_subset_Ici.mpr hab)], - exact is_least_Ici } }, - apply e1.1, - rw set.image_pair, - exact set.mem_insert _ _, -end - -end scott_continuous diff --git a/src/order/liminf_limsup.lean b/src/order/liminf_limsup.lean index adafe31cd2008..ececbb65f3515 100644 --- a/src/order/liminf_limsup.lean +++ b/src/order/liminf_limsup.lean @@ -122,7 +122,7 @@ lemma not_is_bounded_under_of_tendsto_at_bot [preorder β] [no_min_order β] {f ¬ is_bounded_under (≥) l f := @not_is_bounded_under_of_tendsto_at_top α βᵒᵈ _ _ _ _ _ hf -lemma is_bounded_under.bdd_above_range_of_cofinite [semilattice_sup β] {f : α → β} +lemma is_bounded_under.bdd_above_range_of_cofinite [preorder β] [is_directed β (≤)] {f : α → β} (hf : is_bounded_under (≤) cofinite f) : bdd_above (range f) := begin rcases hf with ⟨b, hb⟩, @@ -131,17 +131,17 @@ begin exact ⟨⟨b, ball_image_iff.2 $ λ x, id⟩, (hb.image f).bdd_above⟩ end -lemma is_bounded_under.bdd_below_range_of_cofinite [semilattice_inf β] {f : α → β} +lemma is_bounded_under.bdd_below_range_of_cofinite [preorder β] [is_directed β (≥)] {f : α → β} (hf : is_bounded_under (≥) cofinite f) : bdd_below (range f) := -@is_bounded_under.bdd_above_range_of_cofinite α βᵒᵈ _ _ hf +@is_bounded_under.bdd_above_range_of_cofinite α βᵒᵈ _ _ _ hf -lemma is_bounded_under.bdd_above_range [semilattice_sup β] {f : ℕ → β} +lemma is_bounded_under.bdd_above_range [preorder β] [is_directed β (≤)] {f : ℕ → β} (hf : is_bounded_under (≤) at_top f) : bdd_above (range f) := 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 : ℕ → β} +lemma is_bounded_under.bdd_below_range [preorder β] [is_directed β (≥)] {f : ℕ → β} (hf : is_bounded_under (≥) at_top f) : bdd_below (range f) := -@is_bounded_under.bdd_above_range βᵒᵈ _ _ 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 `≥`. @@ -198,6 +198,31 @@ lemma is_cobounded.mono (h : f ≤ g) : f.is_cobounded r → g.is_cobounded r end relation +section nonempty +variables [preorder α] [nonempty α] {f : filter β} {u : β → α} + +lemma is_bounded_le_at_bot : (at_bot : filter α).is_bounded (≤) := +‹nonempty α›.elim $ λ a, ⟨a, eventually_le_at_bot _⟩ + +lemma is_bounded_ge_at_top : (at_top : filter α).is_bounded (≥) := +‹nonempty α›.elim $ λ a, ⟨a, eventually_ge_at_top _⟩ + +lemma tendsto.is_bounded_under_le_at_bot (h : tendsto u f at_bot) : f.is_bounded_under (≤) u := +is_bounded_le_at_bot.mono h + +lemma tendsto.is_bounded_under_ge_at_top (h : tendsto u f at_top) : f.is_bounded_under (≥) u := +is_bounded_ge_at_top.mono h + +lemma bdd_above_range_of_tendsto_at_top_at_bot [is_directed α (≤)] {u : ℕ → α} + (hx : tendsto u at_top at_bot) : bdd_above (set.range u) := +hx.is_bounded_under_le_at_bot.bdd_above_range + +lemma bdd_below_range_of_tendsto_at_top_at_top [is_directed α (≥)] {u : ℕ → α} + (hx : tendsto u at_top at_top) : bdd_below (set.range u) := +hx.is_bounded_under_ge_at_top.bdd_below_range + +end nonempty + lemma is_cobounded_le_of_bot [preorder α] [order_bot α] {f : filter α} : f.is_cobounded (≤) := ⟨⊥, assume a h, bot_le⟩ @@ -955,6 +980,15 @@ lemma frequently_lt_of_liminf_lt {α β} [conditionally_complete_linear_order β ∃ᶠ x in f, u x < b := @frequently_lt_of_lt_limsup _ βᵒᵈ _ f u b hu h +variables [conditionally_complete_linear_order α] {f : filter α} {b : α} + +lemma lt_mem_sets_of_Limsup_lt (h : f.is_bounded (≤)) (l : f.Limsup < b) : ∀ᶠ a in f, a < b := +let ⟨c, (h : ∀ᶠ a in f, a ≤ c), hcb⟩ := exists_lt_of_cInf_lt h l in +mem_of_superset h $ λ a, hcb.trans_le' + +lemma gt_mem_sets_of_Liminf_gt : f.is_bounded (≥) → b < f.Liminf → ∀ᶠ a in f, b < a := +@lt_mem_sets_of_Limsup_lt αᵒᵈ _ _ _ + end conditionally_complete_linear_order end filter diff --git a/src/topology/algebra/order/liminf_limsup.lean b/src/topology/algebra/order/liminf_limsup.lean index 4e819a7cdbfb6..c16bae4dda725 100644 --- a/src/topology/algebra/order/liminf_limsup.lean +++ b/src/topology/algebra/order/liminf_limsup.lean @@ -51,19 +51,6 @@ lemma filter.tendsto.is_cobounded_under_ge {f : filter β} {u : β → α} {a : [ne_bot f] (h : tendsto u f (𝓝 a)) : f.is_cobounded_under (≥) u := h.is_bounded_under_le.is_cobounded_flip -lemma is_bounded_le_at_bot (α : Type*) [hα : nonempty α] [preorder α] : - (at_bot : filter α).is_bounded (≤) := -is_bounded_iff.2 ⟨set.Iic hα.some, mem_at_bot _, hα.some, λ x hx, hx⟩ - -lemma filter.tendsto.is_bounded_under_le_at_bot {α : Type*} [nonempty α] [preorder α] - {f : filter β} {u : β → α} (h : tendsto u f at_bot) : - f.is_bounded_under (≤) u := -(is_bounded_le_at_bot α).mono h - -lemma bdd_above_range_of_tendsto_at_top_at_bot {α : Type*} [nonempty α] [semilattice_sup α] - {u : ℕ → α} (hx : tendsto u at_top at_bot) : bdd_above (set.range u) := -(filter.tendsto.is_bounded_under_le_at_bot hx).bdd_above_range - end order_closed_topology section order_closed_topology @@ -90,33 +77,11 @@ lemma filter.tendsto.is_cobounded_under_le {f : filter β} {u : β → α} {a : [ne_bot f] (h : tendsto u f (𝓝 a)) : f.is_cobounded_under (≤) u := h.is_bounded_under_ge.is_cobounded_flip -lemma is_bounded_ge_at_top (α : Type*) [hα : nonempty α] [preorder α] : - (at_top : filter α).is_bounded (≥) := -is_bounded_le_at_bot αᵒᵈ - -lemma filter.tendsto.is_bounded_under_ge_at_top {α : Type*} [nonempty α] [preorder α] - {f : filter β} {u : β → α} (h : tendsto u f at_top) : - f.is_bounded_under (≥) u := -(is_bounded_ge_at_top α).mono h - -lemma bdd_below_range_of_tendsto_at_top_at_top {α : Type*} [nonempty α] [semilattice_inf α] - {u : ℕ → α} (hx : tendsto u at_top at_top) : bdd_below (set.range u) := -(filter.tendsto.is_bounded_under_ge_at_top hx).bdd_below_range - end order_closed_topology section conditionally_complete_linear_order variables [conditionally_complete_linear_order α] -theorem lt_mem_sets_of_Limsup_lt {f : filter α} {b} (h : f.is_bounded (≤)) (l : f.Limsup < b) : - ∀ᶠ a in f, a < b := -let ⟨c, (h : ∀ᶠ a in f, a ≤ c), hcb⟩ := exists_lt_of_cInf_lt h l in -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 αᵒᵈ _ - variables [topological_space α] [order_topology α] /-- If the liminf and the limsup of a filter coincide, then this filter converges to From 442a83d738cb208d3600056c489be16900ba701d Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 31 Aug 2023 08:51:42 +0000 Subject: [PATCH 1282/1291] feat(data/finset/lattice): `sup'`/`inf'` lemmas (#18989) Match (most of) the lemmas between `finset.sup`/`finset.inf` and `finset.sup'`/`finset.inf'`. Also golf two proofs using `eq_of_forall_ge_iff` to make sure both APIs prove their lemmas in as closely as possible a way. Also define `finset.nontrivial` to match `set.nontrivial`. --- src/data/dfinsupp/multiset.lean | 5 +- src/data/finset/basic.lean | 80 +++++++------ src/data/finset/lattice.lean | 107 +++++++++++++++--- src/data/finset/locally_finite.lean | 6 +- .../gauss_eisenstein_lemmas.lean | 2 +- 5 files changed, 140 insertions(+), 60 deletions(-) diff --git a/src/data/dfinsupp/multiset.lean b/src/data/dfinsupp/multiset.lean index de2ed3c8f1031..9b20c21615ef9 100644 --- a/src/data/dfinsupp/multiset.lean +++ b/src/data/dfinsupp/multiset.lean @@ -53,9 +53,8 @@ def to_dfinsupp : multiset α →+ Π₀ a : α, ℕ := @[simp] lemma to_dfinsupp_apply (s : multiset α) (a : α) : s.to_dfinsupp a = s.count a := rfl -@[simp] lemma to_dfinsupp_support (s : multiset α) : - s.to_dfinsupp.support = s.to_finset := -(finset.filter_eq_self _).mpr (λ x hx, count_ne_zero.mpr $ multiset.mem_to_finset.1 hx) +@[simp] lemma to_dfinsupp_support (s : multiset α) : s.to_dfinsupp.support = s.to_finset := +finset.filter_true_of_mem $ λ x hx, count_ne_zero.mpr $ multiset.mem_to_finset.1 hx @[simp] lemma to_dfinsupp_replicate (a : α) (n : ℕ) : to_dfinsupp (multiset.replicate n a) = dfinsupp.single a n := diff --git a/src/data/finset/basic.lean b/src/data/finset/basic.lean index ae0db3da0b9b8..9bb6c39eb6a06 100644 --- a/src/data/finset/basic.lean +++ b/src/data/finset/basic.lean @@ -536,14 +536,27 @@ by rw [←coe_ssubset, coe_singleton, set.ssubset_singleton_iff, coe_eq_empty] lemma eq_empty_of_ssubset_singleton {s : finset α} {x : α} (hs : s ⊂ {x}) : s = ∅ := ssubset_singleton_iff.1 hs -lemma eq_singleton_or_nontrivial (ha : a ∈ s) : s = {a} ∨ (s : set α).nontrivial := +/-- A finset is nontrivial if it has at least two elements. -/ +@[reducible] protected def nontrivial (s : finset α) : Prop := (s : set α).nontrivial + +@[simp] lemma not_nontrivial_empty : ¬ (∅ : finset α).nontrivial := by simp [finset.nontrivial] + +@[simp] lemma not_nontrivial_singleton : ¬ ({a} : finset α).nontrivial := +by simp [finset.nontrivial] + +lemma nontrivial.ne_singleton (hs : s.nontrivial) : s ≠ {a} := +by { rintro rfl, exact not_nontrivial_singleton hs } + +lemma eq_singleton_or_nontrivial (ha : a ∈ s) : s = {a} ∨ s.nontrivial := by { rw ←coe_eq_singleton, exact set.eq_singleton_or_nontrivial ha } -lemma nonempty.exists_eq_singleton_or_nontrivial : - s.nonempty → (∃ a, s = {a}) ∨ (s : set α).nontrivial := +lemma nontrivial_iff_ne_singleton (ha : a ∈ s) : s.nontrivial ↔ s ≠ {a} := +⟨nontrivial.ne_singleton, (eq_singleton_or_nontrivial ha).resolve_left⟩ + +lemma nonempty.exists_eq_singleton_or_nontrivial : s.nonempty → (∃ a, s = {a}) ∨ s.nontrivial := λ ⟨a, ha⟩, (eq_singleton_or_nontrivial ha).imp_left $ exists.intro a -instance [nonempty α] : nontrivial (finset α) := +instance nontrivial' [nonempty α] : nontrivial (finset α) := ‹nonempty α›.elim $ λ a, ⟨⟨{a}, ∅, singleton_ne_empty _⟩⟩ instance [is_empty α] : unique (finset α) := @@ -577,6 +590,8 @@ lemma forall_of_forall_cons {p : α → Prop} {h : a ∉ s} (H : ∀ x, x ∈ co @[simp] lemma mk_cons {s : multiset α} (h : (a ::ₘ s).nodup) : (⟨a ::ₘ s, h⟩ : finset α) = cons a ⟨s, (nodup_cons.1 h).2⟩ (nodup_cons.1 h).1 := rfl +@[simp] lemma cons_empty (a : α) : cons a ∅ (not_mem_empty _) = {a} := rfl + @[simp] lemma nonempty_cons (h : a ∉ s) : (cons a s h).nonempty := ⟨a, mem_cons.2 $ or.inl rfl⟩ @[simp] lemma nonempty_mk {m : multiset α} {hm} : (⟨m, hm⟩ : finset α).nonempty ↔ m ≠ 0 := @@ -1379,7 +1394,7 @@ lemma inter_sdiff (s t u : finset α) : s ∩ (t \ u) = s ∩ t \ u := by { ext @[simp] lemma sdiff_inter_self (s₁ s₂ : finset α) : (s₂ \ s₁) ∩ s₁ = ∅ := inf_sdiff_self_left -@[simp] lemma sdiff_self (s₁ : finset α) : s₁ \ s₁ = ∅ := sdiff_self +@[simp] protected lemma sdiff_self (s₁ : finset α) : s₁ \ s₁ = ∅ := sdiff_self lemma sdiff_inter_distrib_right (s t u : finset α) : s \ (t ∩ u) = (s \ t) ∪ (s \ u) := sdiff_inf @@ -1531,7 +1546,7 @@ by rw [←sdiff_singleton_eq_erase, sdiff_sdiff_eq_sdiff_union (singleton_subset union_comm] lemma sdiff_erase_self (ha : a ∈ s) : s \ s.erase a = {a} := -by rw [sdiff_erase ha, sdiff_self, insert_emptyc_eq] +by rw [sdiff_erase ha, finset.sdiff_self, insert_emptyc_eq] lemma sdiff_sdiff_self_left (s t : finset α) : s \ (s \ t) = s ∩ t := sdiff_sdiff_right_self @@ -1588,10 +1603,10 @@ theorem sizeof_lt_sizeof_of_mem [has_sizeof α] {x : α} {s : finset α} (hx : x @[simp] theorem attach_empty : attach (∅ : finset α) = ∅ := rfl -@[simp] lemma attach_nonempty_iff (s : finset α) : s.attach.nonempty ↔ s.nonempty := +@[simp] lemma attach_nonempty_iff {s : finset α} : s.attach.nonempty ↔ s.nonempty := by simp [finset.nonempty] -@[simp] lemma attach_eq_empty_iff (s : finset α) : s.attach = ∅ ↔ s = ∅ := +@[simp] lemma attach_eq_empty_iff {s : finset α} : s.attach = ∅ ↔ s = ∅ := by simpa [eq_empty_iff_forall_not_mem] /-! ### piecewise -/ @@ -1747,7 +1762,7 @@ end decidable_pi_exists /-! ### filter -/ section filter -variables (p q : α → Prop) [decidable_pred p] [decidable_pred q] +variables (p q : α → Prop) [decidable_pred p] [decidable_pred q] {s : finset α} /-- `filter p s` is the set of elements of `s` that satisfy `p`. -/ def filter (s : finset α) : finset α := ⟨_, s.2.filter p⟩ @@ -1772,38 +1787,34 @@ variable (p) theorem filter_filter (s : finset α) : (s.filter p).filter q = s.filter (λa, p a ∧ q a) := ext $ assume a, by simp only [mem_filter, and_comm, and.left_comm] -lemma filter_true {s : finset α} [h : decidable_pred (λ _, true)] : - @finset.filter α (λ _, true) h s = s := -by ext; simp +lemma filter_comm (s : finset α) : (s.filter p).filter q = (s.filter q).filter p := +by simp_rw [filter_filter, and_comm] + +/- We can simplify an application of filter where the decidability is inferred in "the wrong way" -/ +@[simp] lemma filter_congr_decidable (s : finset α) (p : α → Prop) (h : decidable_pred p) + [decidable_pred p] : @filter α p h s = s.filter p := +by congr -@[simp] theorem filter_false {h} (s : finset α) : @filter α (λa, false) h s = ∅ := -ext $ assume a, by simp only [mem_filter, and_false]; refl +lemma filter_true {h} (s : finset α) : @filter α (λ a, true) h s = s := by ext; simp +lemma filter_false {h} (s : finset α) : @filter α (λ a, false) h s = ∅ := by ext; simp variables {p q} -lemma filter_eq_self (s : finset α) : - s.filter p = s ↔ ∀ x ∈ s, p x := -by simp [finset.ext_iff] +lemma filter_eq_self : s.filter p = s ↔ ∀ ⦃x⦄, x ∈ s → p x := by simp [finset.ext_iff] +lemma filter_eq_empty_iff : s.filter p = ∅ ↔ ∀ ⦃x⦄, x ∈ s → ¬ p x := by simp [finset.ext_iff] + +lemma filter_nonempty_iff {s : finset α} : (s.filter p).nonempty ↔ ∃ a ∈ s, p a := +by simp only [nonempty_iff_ne_empty, ne.def, filter_eq_empty_iff, not_not, not_forall] /-- If all elements of a `finset` satisfy the predicate `p`, `s.filter p` is `s`. -/ -@[simp] lemma filter_true_of_mem {s : finset α} (h : ∀ x ∈ s, p x) : s.filter p = s := -(filter_eq_self s).mpr h +@[simp] lemma filter_true_of_mem (h : ∀ x ∈ s, p x) : s.filter p = s := filter_eq_self.2 h /-- If all elements of a `finset` fail to satisfy the predicate `p`, `s.filter p` is `∅`. -/ -lemma filter_false_of_mem {s : finset α} (h : ∀ x ∈ s, ¬ p x) : s.filter p = ∅ := -eq_empty_of_forall_not_mem (by simpa) +@[simp] lemma filter_false_of_mem (h : ∀ x ∈ s, ¬ p x) : s.filter p = ∅ := filter_eq_empty_iff.2 h -lemma filter_eq_empty_iff (s : finset α) : - (s.filter p = ∅) ↔ ∀ x ∈ s, ¬ p x := -begin - refine ⟨_, filter_false_of_mem⟩, - intros hs, - injection hs with hs', - rwa filter_eq_nil at hs' -end - -lemma filter_nonempty_iff {s : finset α} : (s.filter p).nonempty ↔ ∃ a ∈ s, p a := -by simp only [nonempty_iff_ne_empty, ne.def, filter_eq_empty_iff, not_not, not_forall] +@[simp] lemma filter_const (p : Prop) [decidable p] (s : finset α) : + s.filter (λ a, p) = if p then s else ∅ := +by split_ifs; simp [*] lemma filter_congr {s : finset α} (H : ∀ x ∈ s, p x ↔ q x) : filter p s = filter q s := eq_of_veq $ filter_congr H @@ -1944,11 +1955,6 @@ begin { intro x, simp, intros hx hx₂, refine ⟨or.resolve_left (h hx) hx₂, hx₂⟩ } end -/- We can simplify an application of filter where the decidability is inferred in "the wrong way" -/ -@[simp] lemma filter_congr_decidable {α} (s : finset α) (p : α → Prop) (h : decidable_pred p) - [decidable_pred p] : @filter α p h s = s.filter p := -by congr - section classical open_locale classical /-- The following instance allows us to write `{x ∈ s | p x}` for `finset.filter p s`. diff --git a/src/data/finset/lattice.lean b/src/data/finset/lattice.lean index 208cbac617b55..60a235f53593d 100644 --- a/src/data/finset/lattice.lean +++ b/src/data/finset/lattice.lean @@ -55,10 +55,6 @@ fold_map @[simp] lemma sup_singleton {b : β} : ({b} : finset β).sup f = f b := sup_singleton -lemma sup_union [decidable_eq β] : (s₁ ∪ s₂).sup f = s₁.sup f ⊔ s₂.sup f := -finset.induction_on s₁ (by rw [empty_union, sup_empty, bot_sup_eq]) $ λ a s has ih, -by rw [insert_union, sup_insert, sup_insert, ih, sup_assoc] - lemma sup_sup : s.sup (f ⊔ g) = s.sup f ⊔ s.sup g := begin refine finset.cons_induction_on s _ (λ b t _ h, _), @@ -91,6 +87,9 @@ lemma le_sup {b : β} (hb : b ∈ s) : f b ≤ s.sup f := finset.sup_le_iff.1 le lemma le_sup_of_le {b : β} (hb : b ∈ s) (h : a ≤ f b) : a ≤ s.sup f := h.trans $ le_sup hb +lemma sup_union [decidable_eq β] : (s₁ ∪ s₂).sup f = s₁.sup f ⊔ s₂.sup f := +eq_of_forall_ge_iff $ λ c, by simp [or_imp_distrib, forall_and_distrib] + @[simp] lemma sup_bUnion [decidable_eq β] (s : finset γ) (t : γ → finset β) : (s.bUnion t).sup f = s.sup (λ x, (t x).sup f) := eq_of_forall_ge_iff $ λ c, by simp [@forall_swap _ β] @@ -117,11 +116,7 @@ lemma sup_mono (h : s₁ ⊆ s₂) : s₁.sup f ≤ s₂.sup f := finset.sup_le protected 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 - refine eq_of_forall_ge_iff (λ a, _), - simp_rw finset.sup_le_iff, - exact ⟨λ h c hc b hb, h b hb c hc, λ h b hb c hc, h c hc b hb⟩, -end +eq_of_forall_ge_iff $ λ a, by simpa using forall₂_swap @[simp] lemma sup_attach (s : finset β) (f : β → α) : s.attach.sup (λ x, f x) = s.sup f := (s.attach.sup_map (function.embedding.subtype _) f).symm.trans $ congr_arg _ attach_map_val @@ -129,10 +124,7 @@ end /-- See also `finset.product_bUnion`. -/ lemma sup_product_left (s : finset β) (t : finset γ) (f : β × γ → α) : (s ×ˢ t).sup f = s.sup (λ i, t.sup $ λ i', f ⟨i, i'⟩) := -begin - simp only [le_antisymm_iff, finset.sup_le_iff, mem_product, and_imp, prod.forall], - exact ⟨λ b c hb hc, (le_sup hb).trans' $ le_sup hc, λ b hb c hc, le_sup $ mem_product.2 ⟨hb, hc⟩⟩, -end +eq_of_forall_ge_iff $ λ a, by simp [@forall_swap _ γ] lemma sup_product_right (s : finset β) (t : finset γ) (f : β × γ → α) : (s ×ˢ t).sup f = t.sup (λ i', s.sup $ λ i, f ⟨i, i'⟩) := @@ -288,9 +280,6 @@ fold_map @[simp] lemma inf_singleton {b : β} : ({b} : finset β).inf f = f b := inf_singleton -lemma inf_union [decidable_eq β] : (s₁ ∪ s₂).inf f = s₁.inf f ⊓ s₂.inf f := -@sup_union αᵒᵈ _ _ _ _ _ _ _ - lemma inf_inf : s.inf (f ⊓ g) = s.inf f ⊓ s.inf g := @sup_sup αᵒᵈ _ _ _ _ _ _ @@ -301,6 +290,9 @@ by subst hs; exact finset.fold_congr hfg (f : F) (s : finset ι) (g : ι → α) : f (s.inf g) = s.inf (f ∘ g) := finset.cons_induction_on s (map_top f) $ λ i s _ h, by rw [inf_cons, inf_cons, map_inf, h] +lemma inf_union [decidable_eq β] : (s₁ ∪ s₂).inf f = s₁.inf f ⊓ s₂.inf f := +@sup_union αᵒᵈ _ _ _ _ _ _ _ + @[simp] lemma inf_bUnion [decidable_eq β] (s : finset γ) (t : γ → finset β) : (s.bUnion t).inf f = s.inf (λ x, (t x).inf f) := @sup_bUnion αᵒᵈ _ _ _ _ _ _ _ _ @@ -628,11 +620,28 @@ 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) +lemma sup'_union [decidable_eq β] {s₁ s₂ : finset β} (h₁ : s₁.nonempty) (h₂ : s₂.nonempty) + (f : β → α) : + (s₁ ∪ s₂).sup' (h₁.mono $ subset_union_left _ _) f = s₁.sup' h₁ f ⊔ s₂.sup' h₂ f := +eq_of_forall_ge_iff $ λ a, by simp [or_imp_distrib, forall_and_distrib] + 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) := eq_of_forall_ge_iff $ λ c, by simp [@forall_swap _ β] +protected lemma sup'_comm {t : finset γ} (hs : s.nonempty) (ht : t.nonempty) (f : β → γ → α) : + s.sup' hs (λ b, t.sup' ht (f b)) = t.sup' ht (λ c, s.sup' hs $ λ b, f b c) := +eq_of_forall_ge_iff $ λ a, by simpa using forall₂_swap + +lemma sup'_product_left {t : finset γ} (hs : s.nonempty) (ht : t.nonempty) (f : β × γ → α) : + (s ×ˢ t).sup' (hs.product ht) f = s.sup' hs (λ i, t.sup' ht $ λ i', f ⟨i, i'⟩) := +eq_of_forall_ge_iff $ λ a, by simp [@forall_swap _ γ] + +lemma sup'_product_right {t : finset γ} (hs : s.nonempty) (ht : t.nonempty) (f : β × γ → α) : + (s ×ˢ t).sup' (hs.product ht) f = t.sup' ht (λ i', s.sup' hs $ λ i, f ⟨i, i'⟩) := +by rw [sup'_product_left, finset.sup'_comm] + lemma comp_sup'_eq_sup'_comp [semilattice_sup γ] {s : finset β} (H : s.nonempty) {f : β → α} (g : α → γ) (g_sup : ∀ x y, g (x ⊔ y) = g x ⊔ g y) : g (s.sup' H f) = s.sup' H (g ∘ f) := @@ -675,6 +684,15 @@ begin simp only [sup'_le_iff, h₂] { contextual := tt } end +@[simp] lemma _root_.map_finset_sup' [semilattice_sup β] [sup_hom_class F α β] (f : F) + {s : finset ι} (hs) (g : ι → α) : f (s.sup' hs g) = s.sup' hs (f ∘ g) := +by refine hs.cons_induction _ _; intros; simp [*] + +@[simp] lemma sup'_image [decidable_eq β] {s : finset γ} {f : γ → β} (hs : (s.image f).nonempty) + (g : β → α) (hs': s.nonempty := (nonempty.image_iff _).1 hs) : + (s.image f).sup' hs g = s.sup' hs' (g ∘ f) := +by { rw ←with_bot.coe_eq_coe, simp only [coe_sup', sup_image, with_bot.coe_sup] } + @[simp] lemma sup'_map {s : finset γ} {f : γ ↪ β} (g : β → α) (hs : (s.map f).nonempty) (hs': s.nonempty := finset.map_nonempty.mp hs) : (s.map f).sup' hs g = s.sup' hs' (g ∘ f) := @@ -719,11 +737,28 @@ lemma inf'_le_of_le (hb : b ∈ s) (h : f b ≤ a) : s.inf' ⟨b, hb⟩ f ≤ a @[simp] lemma le_inf'_iff : a ≤ s.inf' H f ↔ ∀ b ∈ s, a ≤ f b := @sup'_le_iff αᵒᵈ _ _ _ H f _ +lemma inf'_union [decidable_eq β] {s₁ s₂ : finset β} (h₁ : s₁.nonempty) (h₂ : s₂.nonempty) + (f : β → α) : + (s₁ ∪ s₂).inf' (h₁.mono $ subset_union_left _ _) f = s₁.inf' h₁ f ⊓ s₂.inf' h₂ f := +@sup'_union αᵒᵈ _ _ _ _ _ h₁ h₂ _ + 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 αᵒᵈ _ _ _ _ _ _ Hs _ Ht +protected lemma inf'_comm {t : finset γ} (hs : s.nonempty) (ht : t.nonempty) (f : β → γ → α) : + s.inf' hs (λ b, t.inf' ht (f b)) = t.inf' ht (λ c, s.inf' hs $ λ b, f b c) := +@finset.sup'_comm αᵒᵈ _ _ _ _ _ hs ht _ + +lemma inf'_product_left {t : finset γ} (hs : s.nonempty) (ht : t.nonempty) (f : β × γ → α) : + (s ×ˢ t).inf' (hs.product ht) f = s.inf' hs (λ i, t.inf' ht $ λ i', f ⟨i, i'⟩) := +@sup'_product_left αᵒᵈ _ _ _ _ _ hs ht _ + +lemma inf'_product_right {t : finset γ} (hs : s.nonempty) (ht : t.nonempty) (f : β × γ → α) : + (s ×ˢ t).inf' (hs.product ht) f = t.inf' ht (λ i', s.inf' hs $ λ i, f ⟨i, i'⟩) := +@sup'_product_right αᵒᵈ _ _ _ _ _ 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) := @@ -741,6 +776,15 @@ lemma inf'_mem (s : set α) (w : ∀ x y ∈ s, x ⊓ y ∈ s) s.inf' H f = t.inf' (h₁ ▸ H) g := @sup'_congr αᵒᵈ _ _ _ H _ _ _ h₁ h₂ +@[simp] lemma _root_.map_finset_inf' [semilattice_inf β] [inf_hom_class F α β] (f : F) + {s : finset ι} (hs) (g : ι → α) : f (s.inf' hs g) = s.inf' hs (f ∘ g) := +by refine hs.cons_induction _ _; intros; simp [*] + +@[simp] lemma inf'_image [decidable_eq β] {s : finset γ} {f : γ → β} (hs : (s.image f).nonempty) + (g : β → α) (hs': s.nonempty := (nonempty.image_iff _).1 hs) : + (s.image f).inf' hs g = s.inf' hs' (g ∘ f) := +@sup'_image αᵒᵈ _ _ _ _ _ _ hs _ hs' + @[simp] lemma inf'_map {s : finset γ} {f : γ ↪ β} (g : β → α) (hs : (s.map f).nonempty) (hs': s.nonempty := finset.map_nonempty.mp hs) : (s.map f).inf' hs g = s.inf' hs' (g ∘ f) := @@ -829,6 +873,37 @@ end inf' @[simp] lemma of_dual_inf' [semilattice_sup α] {s : finset ι} (hs : s.nonempty) (f : ι → αᵒᵈ) : of_dual (s.inf' hs f) = s.sup' hs (of_dual ∘ f) := rfl +section distrib_lattice +variables [distrib_lattice α] {s : finset ι} {t : finset κ} (hs : s.nonempty) (ht : t.nonempty) + {f : ι → α} {g : κ → α} {a : α} + +lemma sup'_inf_distrib_left (f : ι → α) (a : α) : a ⊓ s.sup' hs f = s.sup' hs (λ i, a ⊓ f i) := +begin + refine hs.cons_induction (λ i, _) (λ i s hi hs ih, _) , + { simp }, + { simp_rw [sup'_cons hs, inf_sup_left], + rw ih } +end + +lemma sup'_inf_distrib_right (f : ι → α) (a : α) : s.sup' hs f ⊓ a = s.sup' hs (λ i, f i ⊓ a) := +by { rw [inf_comm, sup'_inf_distrib_left], simp_rw inf_comm } + +lemma sup'_inf_sup' (f : ι → α) (g : κ → α) : + s.sup' hs f ⊓ t.sup' ht g = (s ×ˢ t).sup' (hs.product ht) (λ i, f i.1 ⊓ g i.2) := +by simp_rw [finset.sup'_inf_distrib_right, finset.sup'_inf_distrib_left, sup'_product_left] + +lemma inf'_sup_distrib_left (f : ι → α) (a : α) : a ⊔ s.inf' hs f = s.inf' hs (λ i, a ⊔ f i) := +@sup'_inf_distrib_left αᵒᵈ _ _ _ hs _ _ + +lemma inf'_sup_distrib_right (f : ι → α) (a : α) : s.inf' hs f ⊔ a = s.inf' hs (λ i, f i ⊔ a) := +@sup'_inf_distrib_right αᵒᵈ _ _ _ hs _ _ + +lemma inf'_sup_inf' (f : ι → α) (g : κ → α) : + s.inf' hs f ⊔ t.inf' ht g = (s ×ˢ t).inf' (hs.product ht) (λ i, f i.1 ⊔ g i.2) := +@sup'_inf_sup' αᵒᵈ _ _ _ _ _ hs ht _ _ + +end distrib_lattice + section linear_order variables [linear_order α] {s : finset ι} (H : s.nonempty) {f : ι → α} {a : α} diff --git a/src/data/finset/locally_finite.lean b/src/data/finset/locally_finite.lean index f8b6e5b429cc9..7864e60893459 100644 --- a/src/data/finset/locally_finite.lean +++ b/src/data/finset/locally_finite.lean @@ -199,16 +199,16 @@ end lemma Icc_filter_lt_of_lt_right {a b c : α} [decidable_pred (< c)] (h : b < c) : (Icc a b).filter (< c) = Icc a b := -(finset.filter_eq_self _).2 (λ x hx, lt_of_le_of_lt (mem_Icc.1 hx).2 h) +filter_true_of_mem $ λ x hx, (mem_Icc.1 hx).2.trans_lt h lemma Ioc_filter_lt_of_lt_right {a b c : α} [decidable_pred (< c)] (h : b < c) : (Ioc a b).filter (< c) = Ioc a b := -(finset.filter_eq_self _).2 (λ x hx, lt_of_le_of_lt (mem_Ioc.1 hx).2 h) +filter_true_of_mem $ λ x hx, (mem_Ioc.1 hx).2.trans_lt h lemma Iic_filter_lt_of_lt_right {α} [preorder α] [locally_finite_order_bot α] {a c : α} [decidable_pred (< c)] (h : a < c) : (Iic a).filter (< c) = Iic a := -(finset.filter_eq_self _).2 (λ x hx, lt_of_le_of_lt (mem_Iic.1 hx) h) +filter_true_of_mem $ λ x hx, (mem_Iic.1 hx).trans_lt h variables (a b) [fintype α] diff --git a/src/number_theory/legendre_symbol/gauss_eisenstein_lemmas.lean b/src/number_theory/legendre_symbol/gauss_eisenstein_lemmas.lean index 344697e529726..171dcc9523d6f 100644 --- a/src/number_theory/legendre_symbol/gauss_eisenstein_lemmas.lean +++ b/src/number_theory/legendre_symbol/gauss_eisenstein_lemmas.lean @@ -90,7 +90,7 @@ calc (a ^ (p / 2) * (p / 2)! : zmod p) = (λ _ _ _ _ _ _, id) (λ b h _, ⟨b, by simp [-not_le, *] at *⟩) (by intros; split_ifs at *; simp * at *), - by rw [prod_mul_distrib, this]; simp + by rw [prod_mul_distrib, this, prod_const] ... = (-1)^((Ico 1 (p / 2).succ).filter (λ x : ℕ, ¬(a * x : zmod p).val ≤ p / 2)).card * (p / 2)! : by rw [← prod_nat_cast, finset.prod_eq_multiset_prod, From 8818fdefc78642a7e6afcd20be5c184f3c7d9699 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 12 Sep 2023 13:12:04 +0000 Subject: [PATCH 1283/1291] feat(combinatorics/set_family/ahlswede_zhang): Ahlswede-Zhang identity, part I (#18612) The Ahlswede-Zhang identity is a sharpening of the [Lubell-Yamamoto-Meshalkin identity](https://leanprover-community.github.io/mathlib_docs/combinatorics/set_family/lym.html#finset.sum_card_slice_div_choose_le_one), by expliciting the correction term. This PR defines `finset.truncated_sup`/`finset.truncated_inf`, whose cardinalities show up in the correction term. Co-authored-by: Vladimir Ivanov --- .../set_family/ahlswede_zhang.lean | 264 ++++++++++++++++++ src/data/finset/sups.lean | 44 +++ .../gauss_eisenstein_lemmas.lean | 3 +- 3 files changed, 310 insertions(+), 1 deletion(-) create mode 100644 src/combinatorics/set_family/ahlswede_zhang.lean diff --git a/src/combinatorics/set_family/ahlswede_zhang.lean b/src/combinatorics/set_family/ahlswede_zhang.lean new file mode 100644 index 0000000000000..276084c473ce1 --- /dev/null +++ b/src/combinatorics/set_family/ahlswede_zhang.lean @@ -0,0 +1,264 @@ +/- +Copyright (c) 2023 Yaël Dillies, Vladimir Ivanov. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies, Vladimir Ivanov +-/ +import data.finset.sups + +/-! +# The Ahlswede-Zhang identity + +This file proves the Ahlswede-Zhang identity, which is a nontrivial relation between the size of the +"truncated unions" of a set family. It sharpens the Lubell-Yamamoto-Meshalkin inequality +`finset.sum_card_slice_div_choose_le_one`, by making explicit the correction term. + +For a set family `𝒜`, the Ahlswede-Zhang identity states that the sum of +`|⋂ B ∈ 𝒜, B ⊆ A, B|/(|A| * n.choose |A|)` is exactly `1`. + +## Main declarations + +* `finset.truncated_sup`: `s.truncated_sup a` is the supremum of all `b ≤ a` in `𝒜` if there are + some, or `⊤` if there are none. +* `finset.truncated_inf` `s.truncated_inf a` is the infimum of all `b ≥ a` in `𝒜` if there are + some, or `⊥` if there are none. + +## References + +* [R. Ahlswede, Z. Zhang, *An identity in combinatorial extremal theory*](https://doi.org/10.1016/0001-8708(90)90023-G) +* [D. T. Tru, *An AZ-style identity and Bollobás deficiency*](https://doi.org/10.1016/j.jcta.2007.03.005) +-/ + +open_locale finset_family + +namespace finset +variables {α β : Type*} + +/-! ### Truncated supremum, truncated infimum -/ + +section semilattice_sup +variables [semilattice_sup α] [order_top α] [@decidable_rel α (≤)] + [semilattice_sup β] [bounded_order β] [@decidable_rel β (≤)] {s t : finset α} {a b : α} + +private lemma sup_aux : a ∈ lower_closure (s : set α) → (s.filter $ λ b, a ≤ b).nonempty := +λ ⟨b, hb, hab⟩, ⟨b, mem_filter.2 ⟨hb, hab⟩⟩ + +/-- The infimum of the elements of `s` less than `a` if there are some, otherwise `⊤`. -/ +def truncated_sup (s : finset α) (a : α) : α := +if h : a ∈ lower_closure (s : set α) then (s.filter $ λ b, a ≤ b).sup' (sup_aux h) id else ⊤ + +lemma truncated_sup_of_mem (h : a ∈ lower_closure (s : set α)) : + truncated_sup s a = (s.filter $ λ b, a ≤ b).sup' (sup_aux h) id := dif_pos h + +lemma truncated_sup_of_not_mem (h : a ∉ lower_closure (s : set α)) : truncated_sup s a = ⊤ := +dif_neg h + +@[simp] lemma truncated_sup_empty (a : α) : truncated_sup ∅ a = ⊤ := +truncated_sup_of_not_mem $ by simp + +lemma le_truncated_sup : a ≤ truncated_sup s a := +begin + rw truncated_sup, + split_ifs, + { obtain ⟨ℬ, hb, h⟩ := h, + exact h.trans (le_sup' _ $ mem_filter.2 ⟨hb, h⟩) }, + { exact le_top } +end + +lemma map_truncated_sup (e : α ≃o β) (s : finset α) (a : α) : + e (truncated_sup s a) = truncated_sup (s.map e.to_equiv.to_embedding) (e a) := +begin + have : e a ∈ lower_closure (s.map e.to_equiv.to_embedding : set β) + ↔ a ∈ lower_closure (s : set α), + { simp }, + simp_rw [truncated_sup, apply_dite e, map_finset_sup', map_top, this], + congr' with h, + simp only [filter_map, function.comp, equiv.coe_to_embedding, rel_iso.coe_fn_to_equiv, + order_iso.le_iff_le, id.def], + rw sup'_map, -- TODO: Why can't `simp` use `finset.sup'_map`? + simp only [equiv.coe_to_embedding, rel_iso.coe_fn_to_equiv], +end + +variables [decidable_eq α] + +private lemma lower_aux : + a ∈ lower_closure (↑(s ∪ t) : set α) ↔ + a ∈ lower_closure (s : set α) ∨ a ∈ lower_closure (t : set α) := +by rw [coe_union, lower_closure_union, lower_set.mem_sup_iff] + +lemma truncated_sup_union (hs : a ∈ lower_closure (s : set α)) + (ht : a ∈ lower_closure (t : set α)) : + truncated_sup (s ∪ t) a = truncated_sup s a ⊔ truncated_sup t a := +by simpa only [truncated_sup_of_mem, hs, ht, lower_aux.2 (or.inl hs), filter_union] + using sup'_union _ _ _ + +lemma truncated_sup_union_left (hs : a ∈ lower_closure (s : set α)) + (ht : a ∉ lower_closure (t : set α)) : + truncated_sup (s ∪ t) a = truncated_sup s a := +begin + simp only [mem_lower_closure, mem_coe, exists_prop, not_exists, not_and] at ht, + simp only [truncated_sup_of_mem, hs, filter_union, filter_false_of_mem ht, union_empty, + lower_aux.2 (or.inl hs), ht], +end + +lemma truncated_sup_union_right (hs : a ∉ lower_closure (s : set α)) + (ht : a ∈ lower_closure (t : set α)) : + truncated_sup (s ∪ t) a = truncated_sup t a := +by rw [union_comm, truncated_sup_union_left ht hs] + +lemma truncated_sup_union_of_not_mem (hs : a ∉ lower_closure (s : set α)) + (ht : a ∉ lower_closure (t : set α)) : + truncated_sup (s ∪ t) a = ⊤ := +truncated_sup_of_not_mem $ λ h, (lower_aux.1 h).elim hs ht + +end semilattice_sup + +section semilattice_inf +variables [semilattice_inf α] [bounded_order α] [@decidable_rel α (≤)] + [semilattice_inf β] [bounded_order β] [@decidable_rel β (≤)] {s t : finset α} {a : α} + +private lemma inf_aux : a ∈ upper_closure (s : set α) → (s.filter $ λ b, b ≤ a).nonempty := +λ ⟨b, hb, hab⟩, ⟨b, mem_filter.2 ⟨hb, hab⟩⟩ + +/-- The infimum of the elements of `s` less than `a` if there are some, otherwise `⊥`. -/ +def truncated_inf (s : finset α) (a : α) : α := +if h : a ∈ upper_closure (s : set α) then (s.filter $ λ b, b ≤ a).inf' (inf_aux h) id else ⊥ + +lemma truncated_inf_of_mem (h : a ∈ upper_closure (s : set α)) : + truncated_inf s a = (s.filter $ λ b, b ≤ a).inf' (inf_aux h) id := dif_pos h + +lemma truncated_inf_of_not_mem (h : a ∉ upper_closure (s : set α)) : truncated_inf s a = ⊥ := +dif_neg h + +lemma truncated_inf_le (s : finset α) (a : α) : truncated_inf s a ≤ a := +begin + unfold truncated_inf, + split_ifs, + { obtain ⟨ℬ, hb, h⟩ := h, + exact (inf'_le _ $ mem_filter.2 ⟨hb, h⟩).trans h }, + { exact bot_le } +end + +@[simp] lemma truncated_inf_empty (a : α) : truncated_inf ∅ a = ⊥ := +truncated_inf_of_not_mem $ by simp + +lemma map_truncated_inf (e : α ≃o β) (s : finset α) (a : α) : + e (truncated_inf s a) = truncated_inf (s.map e.to_equiv.to_embedding) (e a) := +begin + have : e a ∈ upper_closure (s.map e.to_equiv.to_embedding : set β) + ↔ a ∈ upper_closure (s : set α), + { simp }, + simp_rw [truncated_inf, apply_dite e, map_finset_inf', map_bot, this], + congr' with h, + simp only [filter_map, function.comp, equiv.coe_to_embedding, rel_iso.coe_fn_to_equiv, + order_iso.le_iff_le, id.def], + rw inf'_map, -- TODO: Why can't `simp` use `finset.inf'_map`? + simp only [equiv.coe_to_embedding, rel_iso.coe_fn_to_equiv], +end + +variables [decidable_eq α] + +private lemma upper_aux : + a ∈ upper_closure (↑(s ∪ t) : set α) ↔ + a ∈ upper_closure (s : set α) ∨ a ∈ upper_closure (t : set α) := +by rw [coe_union, upper_closure_union, upper_set.mem_inf_iff] + +lemma truncated_inf_union (hs : a ∈ upper_closure (s : set α)) + (ht : a ∈ upper_closure (t : set α)) : + truncated_inf (s ∪ t) a = truncated_inf s a ⊓ truncated_inf t a := +by simpa only [truncated_inf_of_mem, hs, ht, upper_aux.2 (or.inl hs), filter_union] + using inf'_union _ _ _ + +lemma truncated_inf_union_left (hs : a ∈ upper_closure (s : set α)) + (ht : a ∉ upper_closure (t : set α)) : + truncated_inf (s ∪ t) a = truncated_inf s a := +begin + simp only [mem_upper_closure, mem_coe, exists_prop, not_exists, not_and] at ht, + simp only [truncated_inf_of_mem, hs, filter_union, filter_false_of_mem ht, union_empty, + upper_aux.2 (or.inl hs), ht], +end + +lemma truncated_inf_union_right (hs : a ∉ upper_closure (s : set α)) + (ht : a ∈ upper_closure (t : set α)) : + truncated_inf (s ∪ t) a = truncated_inf t a := +by rw [union_comm, truncated_inf_union_left ht hs] + +lemma truncated_inf_union_of_not_mem (hs : a ∉ upper_closure (s : set α)) + (ht : a ∉ upper_closure (t : set α)) : + truncated_inf (s ∪ t) a = ⊥ := +truncated_inf_of_not_mem $ by { rw [coe_union, upper_closure_union], exact λ h, h.elim hs ht } + +end semilattice_inf + +section distrib_lattice +variables [distrib_lattice α] [bounded_order α] [decidable_eq α] [@decidable_rel α (≤)] + {s t : finset α} {a : α} + +private lemma infs_aux + : a ∈ lower_closure (↑(s ⊼ t) : set α) ↔ a ∈ lower_closure (s : set α) ⊓ lower_closure t := +by rw [coe_infs, lower_closure_infs, lower_set.mem_inf_iff] + +private lemma sups_aux : + a ∈ upper_closure (↑(s ⊻ t) : set α) ↔ a ∈ upper_closure (s : set α) ⊔ upper_closure t := +by rw [coe_sups, upper_closure_sups, upper_set.mem_sup_iff] + +lemma truncated_sup_infs (hs : a ∈ lower_closure (s : set α)) (ht : a ∈ lower_closure (t : set α)) : + truncated_sup (s ⊼ t) a = truncated_sup s a ⊓ truncated_sup t a := +begin + simp only [truncated_sup_of_mem, hs, ht, infs_aux.2 ⟨hs, ht⟩, sup'_inf_sup', filter_infs_ge], + simp_rw ←image_inf_product, + rw sup'_image, + refl, +end + +lemma truncated_inf_sups (hs : a ∈ upper_closure (s : set α)) (ht : a ∈ upper_closure (t : set α)) : + truncated_inf (s ⊻ t) a = truncated_inf s a ⊔ truncated_inf t a := +begin + simp only [truncated_inf_of_mem, hs, ht, sups_aux.2 ⟨hs, ht⟩, inf'_sup_inf', filter_sups_le], + simp_rw ←image_sup_product, + rw inf'_image, + refl, +end + +lemma truncated_sup_infs_of_not_mem (ha : a ∉ lower_closure (s : set α) ⊓ lower_closure t) : + truncated_sup (s ⊼ t) a = ⊤ := +truncated_sup_of_not_mem $ by rwa [coe_infs, lower_closure_infs] + +lemma truncated_inf_sups_of_not_mem (ha : a ∉ upper_closure (s : set α) ⊔ upper_closure t) : + truncated_inf (s ⊻ t) a = ⊥ := +truncated_inf_of_not_mem $ by rwa [coe_sups, upper_closure_sups] + +end distrib_lattice + +section boolean_algebra +variables [boolean_algebra α] [@decidable_rel α (≤)] {s : finset α} {a : α} + +@[simp] lemma compl_truncated_sup (s : finset α) (a : α) : + (truncated_sup s a)ᶜ = truncated_inf (s.map ⟨compl, compl_injective⟩) aᶜ := +map_truncated_sup (order_iso.compl α) _ _ + +@[simp] lemma compl_truncated_inf (s : finset α) (a : α) : + (truncated_inf s a)ᶜ = truncated_sup (s.map ⟨compl, compl_injective⟩) aᶜ := +map_truncated_inf (order_iso.compl α) _ _ + +end boolean_algebra + +variables [decidable_eq α] [fintype α] + +lemma card_truncated_sup_union_add_card_truncated_sup_infs (𝒜 ℬ : finset (finset α)) + (s : finset α) : + (truncated_sup (𝒜 ∪ ℬ) s).card + (truncated_sup (𝒜 ⊼ ℬ) s).card = + (truncated_sup 𝒜 s).card + (truncated_sup ℬ s).card := +begin + by_cases h𝒜 : s ∈ lower_closure (𝒜 : set $ finset α); + by_cases hℬ : s ∈ lower_closure (ℬ : set $ finset α), + { rw [truncated_sup_union h𝒜 hℬ, truncated_sup_infs h𝒜 hℬ], + exact card_union_add_card_inter _ _ }, + { rw [truncated_sup_union_left h𝒜 hℬ, truncated_sup_of_not_mem hℬ, + truncated_sup_infs_of_not_mem (λ h, hℬ h.2)] }, + { rw [truncated_sup_union_right h𝒜 hℬ, truncated_sup_of_not_mem h𝒜, + truncated_sup_infs_of_not_mem (λ h, h𝒜 h.1), add_comm] }, + { rw [truncated_sup_of_not_mem h𝒜, truncated_sup_of_not_mem hℬ, + truncated_sup_union_of_not_mem h𝒜 hℬ, truncated_sup_infs_of_not_mem (λ h, h𝒜 h.1)] } +end + +end finset diff --git a/src/data/finset/sups.lean b/src/data/finset/sups.lean index 0c2c69a056427..b94a8777cedc4 100644 --- a/src/data/finset/sups.lean +++ b/src/data/finset/sups.lean @@ -36,6 +36,20 @@ We define the following notation in locale `finset_family`: open function open_locale set_family +-- TODO: Is there a better spot for those two instances? +namespace finset +variables {α : Type*} [preorder α] {s t : set α} {a : α} + +instance decidable_pred_mem_upper_closure (s : finset α) [@decidable_rel α (≤)] : + decidable_pred (∈ upper_closure (s : set α)) := +λ _, finset.decidable_dexists_finset + +instance decidable_pred_mem_lower_closure (s : finset α) [@decidable_rel α (≤)] : + decidable_pred (∈ lower_closure (s : set α)) := +λ _, finset.decidable_dexists_finset + +end finset + variables {α : Type*} [decidable_eq α] namespace finset @@ -116,6 +130,21 @@ lemma sups_right_comm : (s ⊻ t) ⊻ u = (s ⊻ u) ⊻ t := image₂_right_comm lemma sups_sups_sups_comm : (s ⊻ t) ⊻ (u ⊻ v) = (s ⊻ u) ⊻ (t ⊻ v) := image₂_image₂_image₂_comm sup_sup_sup_comm +variables [@decidable_rel α (≤)] + +lemma filter_sups_le (s t : finset α) (a : α) : + (s ⊻ t).filter (λ b, b ≤ a) = s.filter (λ b, b ≤ a) ⊻ t.filter (λ b, b ≤ a) := +begin + ext b, + simp only [mem_filter, mem_sups], + split, + { rintro ⟨⟨b, hb, c, hc, rfl⟩, ha⟩, + rw sup_le_iff at ha, + exact ⟨_, ⟨hb, ha.1⟩, _, ⟨hc, ha.2⟩, rfl⟩ }, + { rintro ⟨b, hb, c, hc, _, rfl⟩, + exact ⟨⟨_, hb.1, _, hc.1, rfl⟩, sup_le hb.2 hc.2⟩ } +end + end sups section infs @@ -195,6 +224,21 @@ lemma infs_right_comm : (s ⊼ t) ⊼ u = (s ⊼ u) ⊼ t := image₂_right_comm lemma infs_infs_infs_comm : (s ⊼ t) ⊼ (u ⊼ v) = (s ⊼ u) ⊼ (t ⊼ v) := image₂_image₂_image₂_comm inf_inf_inf_comm +variables [@decidable_rel α (≤)] + +lemma filter_infs_ge (s t : finset α) (a : α) : + (s ⊼ t).filter (λ b, a ≤ b) = s.filter (λ b, a ≤ b) ⊼ t.filter (λ b, a ≤ b) := +begin + ext b, + simp only [mem_filter, mem_infs], + split, + { rintro ⟨⟨b, hb, c, hc, rfl⟩, ha⟩, + rw le_inf_iff at ha, + exact ⟨_, ⟨hb, ha.1⟩, _, ⟨hc, ha.2⟩, rfl⟩ }, + { rintro ⟨b, hb, c, hc, _, rfl⟩, + exact ⟨⟨_, hb.1, _, hc.1, rfl⟩, le_inf hb.2 hc.2⟩ } +end + end infs open_locale finset_family diff --git a/src/number_theory/legendre_symbol/gauss_eisenstein_lemmas.lean b/src/number_theory/legendre_symbol/gauss_eisenstein_lemmas.lean index 171dcc9523d6f..293f6f86796bb 100644 --- a/src/number_theory/legendre_symbol/gauss_eisenstein_lemmas.lean +++ b/src/number_theory/legendre_symbol/gauss_eisenstein_lemmas.lean @@ -114,13 +114,14 @@ lemma gauss_lemma {p : ℕ} [fact p.prime] {a : ℤ} (hp : p ≠ 2) (ha0 : (a : (λ x : ℕ, p / 2 < (a * x : zmod p).val)).card := begin haveI hp' : fact (p % 2 = 1) := ⟨nat.prime.mod_two_eq_one_iff_ne_two.mpr hp⟩, + haveI : fact (2 < p) := ⟨hp.lt_of_le' (fact.out p.prime).two_le⟩, have : (legendre_sym p a : zmod p) = (((-1)^((Ico 1 (p / 2).succ).filter (λ x : ℕ, p / 2 < (a * x : zmod p).val)).card : ℤ) : zmod p) := by { rw [legendre_sym.eq_pow, gauss_lemma_aux p ha0]; simp }, cases legendre_sym.eq_one_or_neg_one p ha0; cases neg_one_pow_eq_or ℤ ((Ico 1 (p / 2).succ).filter (λ x : ℕ, p / 2 < (a * x : zmod p).val)).card; - simp [*, ne_neg_self p one_ne_zero, (ne_neg_self p one_ne_zero).symm] at * + simp only [*, neg_one_ne_one, neg_one_ne_one.symm, algebra_map.coe_one, int.cast_neg] at *, end private lemma eisenstein_lemma_aux₁ (p : ℕ) [fact p.prime] [hp2 : fact (p % 2 = 1)] From 001ffdc42920050657fd45bd2b8bfbec8eaaeb29 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 12 Sep 2023 23:46:39 +0000 Subject: [PATCH 1284/1291] feat(probability/independence): Independence of singletons (#18506) Characterisation of independence in terms of measure distributing over finite intersections, and lemmas connecting the different concepts of independence. Also add supporting `measurable_space` and `set.preimage` lemmas --- src/data/set/basic.lean | 3 + src/data/set/image.lean | 10 +++ src/measure_theory/measurable_space.lean | 93 ++++++++++++++++++++ src/measure_theory/measurable_space_def.lean | 6 +- src/probability/independence/basic.lean | 39 +++++++- 5 files changed, 146 insertions(+), 5 deletions(-) diff --git a/src/data/set/basic.lean b/src/data/set/basic.lean index 9ac1590fab649..065bb077cb029 100644 --- a/src/data/set/basic.lean +++ b/src/data/set/basic.lean @@ -1906,3 +1906,6 @@ lemma subset_right_of_subset_union (h : s ⊆ t ∪ u) (hab : disjoint s t) : s hab.left_le_of_le_sup_left h end disjoint + +@[simp] lemma Prop.compl_singleton (p : Prop) : ({p}ᶜ : set Prop) = {¬ p} := +ext $ λ q, by simpa [@iff.comm q] using not_iff diff --git a/src/data/set/image.lean b/src/data/set/image.lean index 4e942d43a7266..9b00e12e92842 100644 --- a/src/data/set/image.lean +++ b/src/data/set/image.lean @@ -71,6 +71,9 @@ theorem subset_preimage_univ {s : set α} : s ⊆ f ⁻¹' univ := subset_univ _ @[simp] theorem preimage_diff (f : α → β) (s t : set β) : f ⁻¹' (s \ t) = f ⁻¹' s \ f ⁻¹' t := rfl +@[simp] lemma preimage_symm_diff (f : α → β) (s t : set β) : + f ⁻¹' (s ∆ t) = (f ⁻¹' s) ∆ (f ⁻¹' t) := rfl + @[simp] theorem preimage_ite (f : α → β) (s t₁ t₂ : set β) : f ⁻¹' (s.ite t₁ t₂) = (f ⁻¹' s).ite (f ⁻¹' t₁) (f ⁻¹' t₂) := rfl @@ -120,6 +123,10 @@ lemma nonempty_of_nonempty_preimage {s : set β} {f : α → β} (hf : (f ⁻¹' s.nonempty := let ⟨x, hx⟩ := hf in ⟨f x, hx⟩ +@[simp] lemma preimage_singleton_true (p : α → Prop) : p ⁻¹' {true} = {a | p a} := by { ext, simp } +@[simp] lemma preimage_singleton_false (p : α → Prop) : p ⁻¹' {false} = {a | ¬ p a} := +by { ext, simp } + lemma preimage_subtype_coe_eq_compl {α : Type*} {s u v : set α} (hsuv : s ⊆ u ∪ v) (H : s ∩ (u ∩ v) = ∅) : (coe : s → α) ⁻¹' u = (coe ⁻¹' v)ᶜ := begin @@ -1003,6 +1010,9 @@ lemma left_inverse.preimage_preimage {g : β → α} (h : left_inverse g f) (s : f ⁻¹' (g ⁻¹' s) = s := by rw [← preimage_comp, h.comp_eq_id, preimage_id] +protected lemma involutive.preimage {f : α → α} (hf : involutive f) : involutive (preimage f) := +hf.right_inverse.preimage_preimage + end function namespace equiv_like diff --git a/src/measure_theory/measurable_space.lean b/src/measure_theory/measurable_space.lean index 98a586b1589a6..74b43d1b7bce7 100644 --- a/src/measure_theory/measurable_space.lean +++ b/src/measure_theory/measurable_space.lean @@ -7,6 +7,7 @@ Authors: Johannes Hölzl, Mario Carneiro import data.prod.tprod import group_theory.coset import logic.equiv.fin +import logic.lemmas import measure_theory.measurable_space_def import order.filter.small_sets import order.liminf_limsup @@ -136,6 +137,12 @@ lemma le_map_comap : m ≤ (m.comap g).map g := (gc_comap_map g).le_u_l _ end functors +@[simp] lemma map_const {m} (b : β) : measurable_space.map (λ a : α, b) m = ⊤ := +eq_top_iff.2 $ by { rintro s hs, by_cases b ∈ s; change measurable_set (preimage _ _); simp [*] } + +@[simp] lemma comap_const {m} (b : β) : measurable_space.comap (λ a : α, b) m = ⊥ := +eq_bot_iff.2 $ by { rintro _ ⟨s, -, rfl⟩, by_cases b ∈ s; simp [*] } + lemma comap_generate_from {f : α → β} {s : set (set β)} : (generate_from s).comap f = generate_from (preimage f '' s) := le_antisymm @@ -291,6 +298,7 @@ section constructions instance : measurable_space empty := ⊤ instance : measurable_space punit := ⊤ -- this also works for `unit` instance : measurable_space bool := ⊤ +instance Prop.measurable_space : measurable_space Prop := ⊤ instance : measurable_space ℕ := ⊤ instance : measurable_space ℤ := ⊤ instance : measurable_space ℚ := ⊤ @@ -298,6 +306,7 @@ instance : measurable_space ℚ := ⊤ instance : measurable_singleton_class empty := ⟨λ _, trivial⟩ instance : measurable_singleton_class punit := ⟨λ _, trivial⟩ instance : measurable_singleton_class bool := ⟨λ _, trivial⟩ +instance Prop.measurable_singleton_class : measurable_singleton_class Prop := ⟨λ _, trivial⟩ instance : measurable_singleton_class ℕ := ⟨λ _, trivial⟩ instance : measurable_singleton_class ℤ := ⟨λ _, trivial⟩ instance : measurable_singleton_class ℚ := ⟨λ _, trivial⟩ @@ -340,6 +349,15 @@ begin exact h, end +lemma measurable_to_prop {f : α → Prop} (h : measurable_set (f⁻¹' {true})) : measurable f := +begin + refine measurable_to_countable' (λ x, _), + by_cases hx : x, + { simpa [hx] using h }, + { simpa only [hx, ←preimage_compl, Prop.compl_singleton, not_true, preimage_singleton_false] + using h.compl } +end + lemma measurable_find_greatest' {p : α → ℕ → Prop} [∀ x, decidable_pred (p x)] {N : ℕ} (hN : ∀ k ≤ N, measurable_set {x | nat.find_greatest (p x) N = k}) : measurable (λ x, nat.find_greatest (p x) N) := @@ -860,8 +878,38 @@ end sum instance {α} {β : α → Type*} [m : Πa, measurable_space (β a)] : measurable_space (sigma β) := ⨅a, (m a).map (sigma.mk a) +section prop +variables {p : α → Prop} + +variables [measurable_space α] + +@[simp] lemma measurable_set_set_of : measurable_set {a | p a} ↔ measurable p := +⟨λ h, measurable_to_prop $ by simpa only [preimage_singleton_true], λ h, + by simpa using h (measurable_set_singleton true)⟩ + +@[simp] lemma measurable_mem : measurable (∈ s) ↔ measurable_set s := measurable_set_set_of.symm + +alias measurable_set_set_of ↔ _ measurable.set_of +alias measurable_mem ↔ _ measurable_set.mem + +end prop end constructions +namespace measurable_space + +/-- The sigma-algebra generated by a single set `s` is `{∅, s, sᶜ, univ}`. -/ +@[simp] lemma generate_from_singleton (s : set α) : + generate_from {s} = measurable_space.comap (∈ s) ⊤ := +begin + classical, + letI : measurable_space α := generate_from {s}, + refine le_antisymm (generate_from_le $ λ t ht, ⟨{true}, trivial, by simp [ht.symm]⟩) _, + rintro _ ⟨u, -, rfl⟩, + exact (show measurable_set s, from generate_measurable.basic _ $ mem_singleton s).mem trivial, +end + +end measurable_space + /-- A map `f : α → β` is called a *measurable embedding* if it is injective, measurable, and sends measurable sets to measurable sets. The latter assumption can be replaced with “`f` has measurable inverse `g : range f → α`”, see `measurable_embedding.measurable_range_splitting`, @@ -1052,6 +1100,9 @@ e.to_equiv.image_eq_preimage s measurable_set (e '' s) ↔ measurable_set s := by rw [image_eq_preimage, measurable_set_preimage] +@[simp] lemma map_eq (e : α ≃ᵐ β) : measurable_space.map e ‹_› = ‹_› := +e.measurable.le_map.antisymm' $ λ s, e.measurable_set_preimage.1 + /-- A measurable equivalence is a measurable embedding. -/ protected lemma measurable_embedding (e : α ≃ᵐ β) : measurable_embedding e := { injective := e.injective, @@ -1277,12 +1328,41 @@ def sum_compl {s : set α} [decidable_pred s] (hs : measurable_set s) : s ⊕ (s measurable_to_fun := by {apply measurable.sum_elim; exact measurable_subtype_coe}, measurable_inv_fun := measurable.dite measurable_inl measurable_inr hs } +/-- Convert a measurable involutive function `f` to a measurable permutation with +`to_fun = inv_fun = f`. See also `function.involutive.to_perm`. -/ +@[simps to_equiv] def of_involutive (f : α → α) (hf : involutive f) (hf' : measurable f) : α ≃ᵐ α := +{ measurable_to_fun := hf', + measurable_inv_fun := hf', + ..hf.to_perm _ } + +@[simp] lemma of_involutive_apply (f : α → α) (hf : involutive f) (hf' : measurable f) (a : α) : + of_involutive f hf hf' a = f a := rfl + +@[simp] lemma of_involutive_symm (f : α → α) (hf : involutive f) (hf' : measurable f) : + (of_involutive f hf hf').symm = of_involutive f hf hf' := rfl + end measurable_equiv namespace measurable_embedding variables [measurable_space α] [measurable_space β] [measurable_space γ] {f : α → β} {g : β → α} +@[simp] lemma comap_eq (hf : measurable_embedding f) : measurable_space.comap f ‹_› = ‹_› := +hf.measurable.comap_le.antisymm $ λ s h, + ⟨_, hf.measurable_set_image' h, hf.injective.preimage_image _⟩ + +lemma iff_comap_eq : + measurable_embedding f ↔ + injective f ∧ measurable_space.comap f ‹_› = ‹_› ∧ measurable_set (range f) := +⟨λ hf, ⟨hf.injective, hf.comap_eq, hf.measurable_set_range⟩, λ hf, + { injective := hf.1, + measurable := by { rw ←hf.2.1, exact comap_measurable f }, + measurable_set_image' := begin + rw ←hf.2.1, + rintro _ ⟨s, hs, rfl⟩, + simpa only [image_preimage_eq_inter_range] using hs.inter hf.2.2, + end }⟩ + /-- A set is equivalent to its image under a function `f` as measurable spaces, if `f` is a measurable embedding -/ noncomputable def equiv_image (s : set α) (hf : measurable_embedding f) : @@ -1377,6 +1457,19 @@ end end measurable_embedding +lemma measurable_space.comap_compl {m' : measurable_space β} [boolean_algebra β] + (h : measurable (compl : β → β)) (f : α → β) : + measurable_space.comap (λ a, (f a)ᶜ) infer_instance = measurable_space.comap f infer_instance := +begin + rw ←measurable_space.comap_comp, + congr', + exact (measurable_equiv.of_involutive _ compl_involutive h).measurable_embedding.comap_eq, +end + +@[simp] lemma measurable_space.comap_not (p : α → Prop) : + measurable_space.comap (λ a, ¬ p a) infer_instance = measurable_space.comap p infer_instance := +measurable_space.comap_compl (λ _ _, trivial) _ + section countably_generated namespace measurable_space diff --git a/src/measure_theory/measurable_space_def.lean b/src/measure_theory/measurable_space_def.lean index 642af17ecd600..6fb980dffc3f7 100644 --- a/src/measure_theory/measurable_space_def.lean +++ b/src/measure_theory/measurable_space_def.lean @@ -196,7 +196,7 @@ by { cases i, exacts [h₂, h₁] } measurable_set (disjointed f n) := disjointed_rec (λ t i ht, measurable_set.diff ht $ h _) (h n) -@[simp] lemma measurable_set.const (p : Prop) : measurable_set {a : α | p} := +lemma measurable_set.const (p : Prop) : measurable_set {a : α | p} := by { by_cases p; simp [h, measurable_set.empty]; apply measurable_set.univ } /-- Every set has a measurable superset. Declare this as local instance as needed. -/ @@ -377,10 +377,10 @@ begin { exact measurable_set_generate_from ht, }, end -@[simp] lemma generate_from_singleton_empty : generate_from {∅} = (⊥ : measurable_space α) := +lemma generate_from_singleton_empty : generate_from {∅} = (⊥ : measurable_space α) := by { rw [eq_bot_iff, generate_from_le_iff], simp, } -@[simp] lemma generate_from_singleton_univ : generate_from {set.univ} = (⊥ : measurable_space α) := +lemma generate_from_singleton_univ : generate_from {set.univ} = (⊥ : measurable_space α) := by { rw [eq_bot_iff, generate_from_le_iff], simp, } lemma measurable_set_bot_iff {s : set α} : @measurable_set α ⊥ s ↔ (s = ∅ ∨ s = univ) := diff --git a/src/probability/independence/basic.lean b/src/probability/independence/basic.lean index 7223f9b23c01e..03218d5783112 100644 --- a/src/probability/independence/basic.lean +++ b/src/probability/independence/basic.lean @@ -64,7 +64,7 @@ when defining `μ` in the example above, the measurable space used is the last o Part A, Chapter 4. -/ -open measure_theory measurable_space +open measure_theory measurable_space set open_locale big_operators measure_theory ennreal namespace probability_theory @@ -577,7 +577,8 @@ We prove the following equivalences on `indep_set`, for measurable sets `s, t`. * `indep_set s t μ ↔ indep_sets {s} {t} μ`. -/ -variables {s t : set Ω} (S T : set (set Ω)) +variables {s t : set Ω} (S T : set (set Ω)) {π : ι → set (set Ω)} {f : ι → set Ω} + {m0 : measurable_space Ω} {μ : measure Ω} lemma indep_set_iff_indep_sets_singleton {m0 : measurable_space Ω} (hs_meas : measurable_set s) (ht_meas : measurable_set t) @@ -624,6 +625,40 @@ lemma indep_iff_forall_indep_set (m₁ m₂ : measurable_space Ω) {m0 : measura λ h s t hs ht, h s t hs ht s t (measurable_set_generate_from (set.mem_singleton s)) (measurable_set_generate_from (set.mem_singleton t))⟩ +lemma Indep_sets.meas_Inter [fintype ι] (h : Indep_sets π μ) (hf : ∀ i, f i ∈ π i) : + μ (⋂ i, f i) = ∏ i, μ (f i) := +by simp [← h _ (λ i _, hf _)] + +lemma Indep_comap_mem_iff : Indep (λ i, measurable_space.comap (∈ f i) ⊤) μ ↔ Indep_set f μ := +by { simp_rw ←generate_from_singleton, refl } + +alias Indep_comap_mem_iff ↔ _ Indep_set.Indep_comap_mem + +lemma Indep_sets_singleton_iff : + Indep_sets (λ i, {f i}) μ ↔ ∀ t, μ (⋂ i ∈ t, f i) = ∏ i in t, μ (f i) := +forall_congr $ λ t, + ⟨λ h, h $ λ _ _, mem_singleton _, + λ h f hf, begin + refine eq.trans _ (h.trans $ finset.prod_congr rfl $ λ i hi, congr_arg _ $ (hf i hi).symm), + rw Inter₂_congr hf, + end⟩ + +variables [is_probability_measure μ] + +lemma Indep_set_iff_Indep_sets_singleton (hf : ∀ i, measurable_set (f i)) : + Indep_set f μ ↔ Indep_sets (λ i, {f i}) μ := +⟨Indep.Indep_sets $ λ _, rfl, Indep_sets.Indep _ + (λ i, generate_from_le $ by { rintro t (rfl : t = _), exact hf _}) _ + (λ _, is_pi_system.singleton _) $ λ _, rfl⟩ + +lemma Indep_set_iff_measure_Inter_eq_prod (hf : ∀ i, measurable_set (f i)) : + Indep_set f μ ↔ ∀ s, μ (⋂ i ∈ s, f i) = ∏ i in s, μ (f i) := +(Indep_set_iff_Indep_sets_singleton hf).trans Indep_sets_singleton_iff + +lemma Indep_sets.Indep_set_of_mem (hfπ : ∀ i, f i ∈ π i) (hf : ∀ i, measurable_set (f i)) + (hπ : Indep_sets π μ) : Indep_set f μ := +(Indep_set_iff_measure_Inter_eq_prod hf).2 $ λ t, hπ _ $ λ i _, hfπ _ + end indep_set section indep_fun From ce64cd319bb6b3e82f31c2d38e79080d377be451 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 18 Sep 2023 15:31:36 +0000 Subject: [PATCH 1285/1291] feat(topology/algebra/order/liminf_limsup): Eventual boundedness of neighborhoods (#18629) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Generalise `bounded_le_nhds`/`bounded_ge_nhds` using two ad hoc typeclasses. The goal here is to circumvent the fact that the product of order topologies is *not* an order topology, and apply those lemmas to `ℝⁿ`. --- src/order/filter/pi.lean | 9 ++ src/topology/algebra/order/liminf_limsup.lean | 144 ++++++++++++++---- 2 files changed, 123 insertions(+), 30 deletions(-) diff --git a/src/order/filter/pi.lean b/src/order/filter/pi.lean index d069ed0dcd512..1d9167060ab6f 100644 --- a/src/order/filter/pi.lean +++ b/src/order/filter/pi.lean @@ -28,6 +28,7 @@ open_locale classical filter namespace filter variables {ι : Type*} {α : ι → Type*} {f f₁ f₂ : Π i, filter (α i)} {s : Π i, set (α i)} + {p : Π i, α i → Prop} section pi @@ -93,6 +94,14 @@ end I.pi s ∈ pi f ↔ ∀ i ∈ I, s i ∈ f i := ⟨λ h i hi, mem_of_pi_mem_pi h hi, pi_mem_pi hI⟩ +lemma eventually.eval_pi {i : ι} (hf : ∀ᶠ (x : α i) in f i, p i x) : + ∀ᶠ (x : Π (i : ι), α i) in pi f, p i (x i) := +(tendsto_eval_pi _ _).eventually hf + +lemma eventually_pi [finite ι] (hf : ∀ i, ∀ᶠ x in f i, p i x) : + ∀ᶠ (x : Π i, α i) in pi f, ∀ i, p i (x i) := +eventually_all.2 $ λ i, (hf _).eval_pi + lemma has_basis_pi {ι' : ι → Type} {s : Π i, ι' i → set (α i)} {p : Π i, ι' i → Prop} (h : ∀ i, (f i).has_basis (p i) (s i)) : (pi f).has_basis (λ If : set ι × Π i, ι' i, If.1.finite ∧ ∀ i ∈ If.1, p i (If.2 i)) diff --git a/src/topology/algebra/order/liminf_limsup.lean b/src/topology/algebra/order/liminf_limsup.lean index c16bae4dda725..74946971a14e3 100644 --- a/src/topology/algebra/order/liminf_limsup.lean +++ b/src/topology/algebra/order/liminf_limsup.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2017 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Johannes Hölzl, Mario Carneiro, Yury Kudryashov +Authors: Johannes Hölzl, Mario Carneiro, Yury Kudryashov, Yaël Dillies -/ import algebra.big_operators.intervals import algebra.big_operators.order @@ -16,68 +16,152 @@ import topology.order.basic > THIS FILE IS SYNCHRONIZED WITH MATHLIB4. > Any changes to this file require a corresponding PR to mathlib4. + +## Main declarations + +* `bounded_le_nhds_class`: Typeclass stating that neighborhoods are eventually bounded above. +* `bounded_ge_nhds_class`: Typeclass stating that neighborhoods are eventually bounded below. + +## Implementation notes + +The same lemmas are true in `ℝ`, `ℝ × ℝ`, `ι → ℝ`, `euclidean_space ι ℝ`. To avoid code +duplication, we provide an ad hoc axiomatisation of the properties we need. -/ open filter topological_space open_locale topology classical universes u v -variables {α : Type u} {β : Type v} +variables {ι α β R S : Type*} {π : ι → Type*} -section liminf_limsup +/-- Ad hoc typeclass stating that neighborhoods are eventually bounded above. -/ +class bounded_le_nhds_class (α : Type*) [preorder α] [topological_space α] : Prop := +(is_bounded_le_nhds (a : α) : (𝓝 a).is_bounded (≤)) + +/-- Ad hoc typeclass stating that neighborhoods are eventually bounded below. -/ +class bounded_ge_nhds_class (α : Type*) [preorder α] [topological_space α] : Prop := +(is_bounded_ge_nhds (a : α) : (𝓝 a).is_bounded (≥)) + +section preorder +variables [preorder α] [preorder β] [topological_space α] [topological_space β] -section order_closed_topology -variables [semilattice_sup α] [topological_space α] [order_topology α] +section bounded_le_nhds_class +variables [bounded_le_nhds_class α] [bounded_le_nhds_class β] {f : filter ι} {u : ι → α} {a : α} lemma is_bounded_le_nhds (a : α) : (𝓝 a).is_bounded (≤) := -(is_top_or_exists_gt a).elim (λ h, ⟨a, eventually_of_forall h⟩) (λ ⟨b, hb⟩, ⟨b, ge_mem_nhds hb⟩) +bounded_le_nhds_class.is_bounded_le_nhds _ -lemma filter.tendsto.is_bounded_under_le {f : filter β} {u : β → α} {a : α} - (h : tendsto u f (𝓝 a)) : f.is_bounded_under (≤) u := +lemma filter.tendsto.is_bounded_under_le (h : tendsto u f (𝓝 a)) : + f.is_bounded_under (≤) u := (is_bounded_le_nhds a).mono h -lemma filter.tendsto.bdd_above_range_of_cofinite {u : β → α} {a : α} +lemma filter.tendsto.bdd_above_range_of_cofinite [is_directed α (≤)] (h : tendsto u cofinite (𝓝 a)) : bdd_above (set.range u) := h.is_bounded_under_le.bdd_above_range_of_cofinite -lemma filter.tendsto.bdd_above_range {u : ℕ → α} {a : α} - (h : tendsto u at_top (𝓝 a)) : bdd_above (set.range u) := +lemma filter.tendsto.bdd_above_range [is_directed α (≤)] {u : ℕ → α} (h : tendsto u at_top (𝓝 a)) : + bdd_above (set.range u) := h.is_bounded_under_le.bdd_above_range lemma is_cobounded_ge_nhds (a : α) : (𝓝 a).is_cobounded (≥) := (is_bounded_le_nhds a).is_cobounded_flip -lemma filter.tendsto.is_cobounded_under_ge {f : filter β} {u : β → α} {a : α} - [ne_bot f] (h : tendsto u f (𝓝 a)) : f.is_cobounded_under (≥) u := +lemma filter.tendsto.is_cobounded_under_ge [ne_bot f] (h : tendsto u f (𝓝 a)) : + f.is_cobounded_under (≥) u := h.is_bounded_under_le.is_cobounded_flip -end order_closed_topology +instance : bounded_ge_nhds_class αᵒᵈ := ⟨@is_bounded_le_nhds α _ _ _⟩ + +instance : bounded_le_nhds_class (α × β) := +begin + refine ⟨λ x, _⟩, + obtain ⟨a, ha⟩ := is_bounded_le_nhds x.1, + obtain ⟨b, hb⟩ := is_bounded_le_nhds x.2, + rw [←@prod.mk.eta _ _ x, nhds_prod_eq], + exact ⟨(a, b), ha.prod_mk hb⟩, +end + +instance [finite ι] [Π i, preorder (π i)] [Π i, topological_space (π i)] + [Π i, bounded_le_nhds_class (π i)] : bounded_le_nhds_class (Π i, π i) := +begin + refine ⟨λ x, _⟩, + rw nhds_pi, + choose f hf using λ i, is_bounded_le_nhds (x i), + exact ⟨f, eventually_pi hf⟩, +end + +end bounded_le_nhds_class -section order_closed_topology -variables [semilattice_inf α] [topological_space α] [order_topology α] +section bounded_ge_nhds_class +variables [bounded_ge_nhds_class α] [bounded_ge_nhds_class β] {f : filter ι} {u : ι → α} {a : α} -lemma is_bounded_ge_nhds (a : α) : (𝓝 a).is_bounded (≥) := @is_bounded_le_nhds αᵒᵈ _ _ _ a +lemma is_bounded_ge_nhds (a : α) : (𝓝 a).is_bounded (≥) := +bounded_ge_nhds_class.is_bounded_ge_nhds _ -lemma filter.tendsto.is_bounded_under_ge {f : filter β} {u : β → α} {a : α} - (h : tendsto u f (𝓝 a)) : f.is_bounded_under (≥) u := +lemma filter.tendsto.is_bounded_under_ge (h : tendsto u f (𝓝 a)) : + f.is_bounded_under (≥) u := (is_bounded_ge_nhds a).mono h -lemma filter.tendsto.bdd_below_range_of_cofinite {u : β → α} {a : α} +lemma filter.tendsto.bdd_below_range_of_cofinite [is_directed α (≥)] (h : tendsto u cofinite (𝓝 a)) : bdd_below (set.range u) := h.is_bounded_under_ge.bdd_below_range_of_cofinite -lemma filter.tendsto.bdd_below_range {u : ℕ → α} {a : α} - (h : tendsto u at_top (𝓝 a)) : bdd_below (set.range u) := +lemma filter.tendsto.bdd_below_range [is_directed α (≥)] {u : ℕ → α} (h : tendsto u at_top (𝓝 a)) : + bdd_below (set.range u) := h.is_bounded_under_ge.bdd_below_range lemma is_cobounded_le_nhds (a : α) : (𝓝 a).is_cobounded (≤) := (is_bounded_ge_nhds a).is_cobounded_flip -lemma filter.tendsto.is_cobounded_under_le {f : filter β} {u : β → α} {a : α} - [ne_bot f] (h : tendsto u f (𝓝 a)) : f.is_cobounded_under (≤) u := +lemma filter.tendsto.is_cobounded_under_le [ne_bot f] (h : tendsto u f (𝓝 a)) : + f.is_cobounded_under (≤) u := h.is_bounded_under_ge.is_cobounded_flip -end order_closed_topology +instance : bounded_le_nhds_class αᵒᵈ := ⟨@is_bounded_ge_nhds α _ _ _⟩ + +instance : bounded_ge_nhds_class (α × β) := +begin + refine ⟨λ x, _⟩, + obtain ⟨a, ha⟩ := is_bounded_ge_nhds x.1, + obtain ⟨b, hb⟩ := is_bounded_ge_nhds x.2, + rw [←@prod.mk.eta _ _ x, nhds_prod_eq], + exact ⟨(a, b), ha.prod_mk hb⟩, +end + +instance [finite ι] [Π i, preorder (π i)] [Π i, topological_space (π i)] + [Π i, bounded_ge_nhds_class (π i)] : bounded_ge_nhds_class (Π i, π i) := +begin + refine ⟨λ x, _⟩, + rw nhds_pi, + choose f hf using λ i, is_bounded_ge_nhds (x i), + exact ⟨f, eventually_pi hf⟩, +end + +end bounded_ge_nhds_class + +@[priority 100] -- See note [lower instance priority] +instance order_top.to_bounded_le_nhds_class [order_top α] : bounded_le_nhds_class α := +⟨λ a, is_bounded_le_of_top⟩ + +@[priority 100] -- See note [lower instance priority] +instance order_bot.to_bounded_ge_nhds_class [order_bot α] : bounded_ge_nhds_class α := +⟨λ a, is_bounded_ge_of_bot⟩ + +@[priority 100] -- See note [lower instance priority] +instance order_topology.to_bounded_le_nhds_class [is_directed α (≤)] [order_topology α] : + bounded_le_nhds_class α := +⟨λ a, (is_top_or_exists_gt a).elim (λ h, ⟨a, eventually_of_forall h⟩) $ Exists.imp $ λ b, + ge_mem_nhds⟩ + +@[priority 100] -- See note [lower instance priority] +instance order_topology.to_bounded_ge_nhds_class [is_directed α (≥)] [order_topology α] : + bounded_ge_nhds_class α := +⟨λ a, (is_bot_or_exists_lt a).elim (λ h, ⟨a, eventually_of_forall h⟩) $ Exists.imp $ λ b, + le_mem_nhds⟩ + +end preorder + +section liminf_limsup section conditionally_complete_linear_order variables [conditionally_complete_linear_order α] @@ -224,7 +308,7 @@ end liminf_limsup section monotone -variables {ι R S : Type*} {F : filter ι} [ne_bot F] +variables {F : filter ι} [ne_bot F] [complete_linear_order R] [topological_space R] [order_topology R] [complete_linear_order S] [topological_space S] [order_topology S] @@ -334,7 +418,7 @@ open_locale topology open filter set -variables {ι : Type*} {R : Type*} [complete_linear_order R] [topological_space R] [order_topology R] +variables [complete_linear_order R] [topological_space R] [order_topology R] lemma infi_eq_of_forall_le_of_tendsto {x : R} {as : ι → R} (x_le : ∀ i, x ≤ as i) {F : filter ι} [filter.ne_bot F] (as_lim : filter.tendsto as F (𝓝 x)) : @@ -349,7 +433,7 @@ lemma supr_eq_of_forall_le_of_tendsto {x : R} {as : ι → R} (⨆ i, as i) = x := @infi_eq_of_forall_le_of_tendsto ι (order_dual R) _ _ _ x as le_x F _ as_lim -lemma Union_Ici_eq_Ioi_of_lt_of_tendsto {ι : Type*} (x : R) {as : ι → R} (x_lt : ∀ i, x < as i) +lemma Union_Ici_eq_Ioi_of_lt_of_tendsto (x : R) {as : ι → R} (x_lt : ∀ i, x < as i) {F : filter ι} [filter.ne_bot F] (as_lim : filter.tendsto as F (𝓝 x)) : (⋃ (i : ι), Ici (as i)) = Ioi x := begin @@ -361,10 +445,10 @@ begin exact Union_Ici_eq_Ioi_infi obs, end -lemma Union_Iic_eq_Iio_of_lt_of_tendsto {ι : Type*} (x : R) {as : ι → R} (lt_x : ∀ i, as i < x) +lemma Union_Iic_eq_Iio_of_lt_of_tendsto (x : R) {as : ι → R} (lt_x : ∀ i, as i < x) {F : filter ι} [filter.ne_bot F] (as_lim : filter.tendsto as F (𝓝 x)) : (⋃ (i : ι), Iic (as i)) = Iio x := -@Union_Ici_eq_Ioi_of_lt_of_tendsto (order_dual R) _ _ _ ι x as lt_x F _ as_lim +@Union_Ici_eq_Ioi_of_lt_of_tendsto ι Rᵒᵈ _ _ _ _ _ lt_x F _ as_lim end infi_and_supr From c8f305514e0d47dfaa710f5a52f0d21b588e6328 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 9 Oct 2023 14:37:58 +0000 Subject: [PATCH 1286/1291] feat(topology/metric_space): diameter of pointwise zero and addition (#19028) --- src/analysis/normed/group/pointwise.lean | 13 ++++++++++ src/topology/metric_space/antilipschitz.lean | 25 +++++++++++++++++++- src/topology/metric_space/basic.lean | 11 ++++++++- src/topology/metric_space/emetric_space.lean | 4 +++- src/topology/metric_space/lipschitz.lean | 25 ++++++++++++++++++++ 5 files changed, 75 insertions(+), 3 deletions(-) diff --git a/src/analysis/normed/group/pointwise.lean b/src/analysis/normed/group/pointwise.lean index 672dd575ae904..340dc562045b2 100644 --- a/src/analysis/normed/group/pointwise.lean +++ b/src/analysis/normed/group/pointwise.lean @@ -5,6 +5,7 @@ Authors: Sébastien Gouëzel, Yaël Dillies -/ import analysis.normed.group.basic import topology.metric_space.hausdorff_distance +import topology.metric_space.isometric_smul /-! # Properties of pointwise addition of sets in normed groups @@ -24,6 +25,7 @@ variables {E : Type*} section seminormed_group variables [seminormed_group E] {ε δ : ℝ} {s t : set E} {x y : E} +-- note: we can't use `lipschitz_on_with.bounded_image2` here without adding `[isometric_smul E E]` @[to_additive] lemma metric.bounded.mul (hs : bounded s) (ht : bounded t) : bounded (s * t) := begin obtain ⟨Rs, hRs⟩ : ∃ R, ∀ x ∈ s, ‖x‖ ≤ R := hs.exists_norm_le', @@ -33,6 +35,10 @@ begin exact norm_mul_le_of_le (hRs x hx) (hRt y hy), end +@[to_additive] lemma metric.bounded.of_mul (hst : bounded (s * t)) : + bounded s ∨ bounded t := +antilipschitz_with.bounded_of_image2_left _ (λ x, (isometry_mul_right x).antilipschitz) hst + @[to_additive] lemma metric.bounded.inv : bounded s → bounded s⁻¹ := by { simp_rw [bounded_iff_forall_norm_le', ←image_inv, ball_image_iff, norm_inv'], exact id } @@ -55,6 +61,13 @@ eq_of_forall_le_iff $ λ r, by simp_rw [le_inf_edist, ←image_inv, ball_image_i lemma inf_edist_inv_inv (x : E) (s : set E) : inf_edist x⁻¹ s⁻¹ = inf_edist x s := by rw [inf_edist_inv, inv_inv] +@[to_additive] lemma ediam_mul_le (x y : set E) : + emetric.diam (x * y) ≤ emetric.diam x + emetric.diam y := +(lipschitz_on_with.ediam_image2_le (*) _ _ + (λ _ _, (isometry_mul_right _).lipschitz.lipschitz_on_with _) + (λ _ _, (isometry_mul_left _).lipschitz.lipschitz_on_with _)).trans_eq $ + by simp only [ennreal.coe_one, one_mul] + end emetric variables (ε δ s t x y) diff --git a/src/topology/metric_space/antilipschitz.lean b/src/topology/metric_space/antilipschitz.lean index c08fbf9d2ef79..7974ab7cbf9d0 100644 --- a/src/topology/metric_space/antilipschitz.lean +++ b/src/topology/metric_space/antilipschitz.lean @@ -192,7 +192,8 @@ namespace antilipschitz_with open metric -variables [pseudo_metric_space α] [pseudo_metric_space β] {K : ℝ≥0} {f : α → β} +variables [pseudo_metric_space α] [pseudo_metric_space β] [pseudo_metric_space γ] +variables {K : ℝ≥0} {f : α → β} lemma bounded_preimage (hf : antilipschitz_with K f) {s : set β} (hs : bounded s) : @@ -219,6 +220,28 @@ begin exact (hf.image_preimage _).symm end +lemma bounded_of_image2_left (f : α → β → γ) {K₁ : ℝ≥0} + (hf : ∀ b, antilipschitz_with K₁ (λ a, f a b)) + {s : set α} {t : set β} (hst : bounded (set.image2 f s t)) : + bounded s ∨ bounded t := +begin + contrapose! hst, + obtain ⟨b, hb⟩ : t.nonempty := nonempty_of_unbounded hst.2, + have : ¬bounded (set.image2 f s {b}), + { intro h, + apply hst.1, + rw set.image2_singleton_right at h, + replace h := (hf b).bounded_preimage h, + refine h.mono (subset_preimage_image _ _) }, + exact mt (bounded.mono (image2_subset subset.rfl (singleton_subset_iff.mpr hb))) this, +end + +lemma bounded_of_image2_right {f : α → β → γ} {K₂ : ℝ≥0} + (hf : ∀ a, antilipschitz_with K₂ (f a)) + {s : set α} {t : set β} (hst : bounded (set.image2 f s t)) : + bounded s ∨ bounded t := +or.symm $ bounded_of_image2_left (flip f) hf $ image2_swap f s t ▸ hst + end antilipschitz_with lemma lipschitz_with.to_right_inverse [pseudo_emetric_space α] [pseudo_emetric_space β] {K : ℝ≥0} diff --git a/src/topology/metric_space/basic.lean b/src/topology/metric_space/basic.lean index de0baf9a61b0a..76e23c4a7f391 100644 --- a/src/topology/metric_space/basic.lean +++ b/src/topology/metric_space/basic.lean @@ -52,7 +52,7 @@ metric, pseudo_metric, dist open set filter topological_space bornology -open_locale uniformity topology big_operators filter nnreal ennreal +open_locale uniformity topology big_operators filter nnreal ennreal pointwise universes u v w variables {α : Type u} {β : Type v} {X ι : Type*} @@ -2171,6 +2171,13 @@ end @[simp] lemma bounded_empty : bounded (∅ : set α) := ⟨0, by simp⟩ +lemma nonempty_of_unbounded (h : ¬ bounded s) : s.nonempty := +begin + rw nonempty_iff_ne_empty, + rintro rfl, + exact h bounded_empty +end + lemma bounded_iff_mem_bounded : bounded s ↔ ∀ x ∈ s, bounded s := ⟨λ h _ _, h, λ H, s.eq_empty_or_nonempty.elim @@ -2453,6 +2460,8 @@ diam_subsingleton subsingleton_empty @[simp] lemma diam_singleton : diam ({x} : set α) = 0 := diam_subsingleton subsingleton_singleton +@[simp, to_additive] lemma diam_one [has_one α] : diam (1 : set α) = 0 := diam_singleton + -- Does not work as a simp-lemma, since {x, y} reduces to (insert y {x}) lemma diam_pair : diam ({x, y} : set α) = dist x y := by simp only [diam, emetric.diam_pair, dist_edist] diff --git a/src/topology/metric_space/emetric_space.lean b/src/topology/metric_space/emetric_space.lean index 983cb4e4990ce..c30cfc396c9a9 100644 --- a/src/topology/metric_space/emetric_space.lean +++ b/src/topology/metric_space/emetric_space.lean @@ -32,7 +32,7 @@ to `emetric_space` at the end. open set filter classical -open_locale uniformity topology big_operators filter nnreal ennreal +open_locale uniformity topology big_operators filter nnreal ennreal pointwise universes u v w variables {α : Type u} {β : Type v} {X : Type*} @@ -793,6 +793,8 @@ diam_subsingleton subsingleton_empty @[simp] lemma diam_singleton : diam ({x} : set α) = 0 := diam_subsingleton subsingleton_singleton +@[simp, to_additive] lemma diam_one [has_one α] : diam (1 : set α) = 0 := diam_singleton + lemma diam_Union_mem_option {ι : Type*} (o : option ι) (s : ι → set α) : diam (⋃ i ∈ o, s i) = ⨆ i ∈ o, diam (s i) := by cases o; simp diff --git a/src/topology/metric_space/lipschitz.lean b/src/topology/metric_space/lipschitz.lean index 08d7ec3473268..235c668ff913e 100644 --- a/src/topology/metric_space/lipschitz.lean +++ b/src/topology/metric_space/lipschitz.lean @@ -465,6 +465,20 @@ protected lemma comp {g : β → γ} {t : set β} {Kg : ℝ≥0} (hg : lipschitz lipschitz_on_with (Kg * K) (g ∘ f) s := lipschitz_on_with_iff_restrict.mpr $ hg.to_restrict.comp (hf.to_restrict_maps_to hmaps) +lemma ediam_image2_le (f : α → β → γ) {K₁ K₂ : ℝ≥0} + (s : set α) (t : set β) + (hf₁ : ∀ b ∈ t, lipschitz_on_with K₁ (λ a, f a b) s) + (hf₂ : ∀ a ∈ s, lipschitz_on_with K₂ (f a) t) : + emetric.diam (set.image2 f s t) ≤ ↑K₁ * emetric.diam s + ↑K₂ * emetric.diam t := +begin + apply emetric.diam_le, + rintros _ ⟨a₁, b₁, ha₁, hb₁, rfl⟩ _ ⟨a₂, b₂, ha₂, hb₂, rfl⟩, + refine (edist_triangle _ (f a₂ b₁) _).trans _, + exact add_le_add + ((hf₁ b₁ hb₁ ha₁ ha₂).trans $ ennreal.mul_left_mono $ emetric.edist_le_diam_of_mem ha₁ ha₂) + ((hf₂ a₂ ha₂ hb₁ hb₂).trans $ ennreal.mul_left_mono $ emetric.edist_le_diam_of_mem hb₁ hb₂), +end + end emetric section metric @@ -512,6 +526,17 @@ protected lemma iff_le_add_mul {f : α → ℝ} {K : ℝ≥0} : lipschitz_on_with K f s ↔ ∀ (x ∈ s) (y ∈ s), f x ≤ f y + K * dist x y := ⟨lipschitz_on_with.le_add_mul, lipschitz_on_with.of_le_add_mul K⟩ +lemma bounded_image2 (f : α → β → γ) {K₁ K₂ : ℝ≥0} + {s : set α} {t : set β} (hs : metric.bounded s) (ht : metric.bounded t) + (hf₁ : ∀ b ∈ t, lipschitz_on_with K₁ (λ a, f a b) s) + (hf₂ : ∀ a ∈ s, lipschitz_on_with K₂ (f a) t) : + metric.bounded (set.image2 f s t) := +metric.bounded_iff_ediam_ne_top.2 $ ne_top_of_le_ne_top + (ennreal.add_ne_top.mpr ⟨ + ennreal.mul_ne_top ennreal.coe_ne_top hs.ediam_ne_top, + ennreal.mul_ne_top ennreal.coe_ne_top ht.ediam_ne_top⟩) + (ediam_image2_le _ _ _ hf₁ hf₂) + end metric end lipschitz_on_with From 3ac76ec59aac51aceffc1e4888f870497dc0ab64 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 10 Oct 2023 22:55:10 +0000 Subject: [PATCH 1287/1291] doc: Add a warning mentioning Lean 4 to the readme (#19243) Also correct links to point to the lean3 webpages; this means that users clicking them end up on pages which also have scary banners telling them not to use Lean 3. --- README.md | 33 ++++++++++++++++++++------------- 1 file changed, 20 insertions(+), 13 deletions(-) diff --git a/README.md b/README.md index 88b3c18f3ccbe..53d8c3defebf5 100644 --- a/README.md +++ b/README.md @@ -1,28 +1,32 @@ -# Lean mathlib +# Lean 3's mathlib + +> [!WARNING] +> Lean 3 and Mathlib 3 are no longer actively maintained. +> It is strongly recommended that you use [mathlib4](https://github.com/leanprover-community/mathlib4) for Lean 4 instead. ![](https://github.com/leanprover-community/mathlib/workflows/continuous%20integration/badge.svg?branch=master) [![Bors enabled](https://bors.tech/images/badge_small.svg)](https://app.bors.tech/repositories/24316) [![project chat](https://img.shields.io/badge/zulip-join_chat-brightgreen.svg)](https://leanprover.zulipchat.com) [![Gitpod Ready-to-Code](https://img.shields.io/badge/Gitpod-ready--to--code-blue?logo=gitpod)](https://gitpod.io/#https://github.com/leanprover-community/mathlib) -[Mathlib](https://leanprover-community.github.io) is a user maintained library for the [Lean theorem prover](https://leanprover.github.io). +[Mathlib](https://leanprover-community.github.io) is a user maintained library for the [Lean 3 theorem prover](https://github.com/leanprover-community/lean). It contains both programming infrastructure and mathematics, as well as tactics that use the former and allow to develop the latter. ## Installation -You can find detailed instructions to install Lean, mathlib, and supporting tools on [our website](https://leanprover-community.github.io/get_started.html). +You can find detailed instructions to install Lean 3, mathlib 3, and supporting tools on [our website](https://leanprover-community.github.io/lean3/get_started.html). ## Experimenting -Got everything installed? Why not start with the [tutorial project](https://leanprover-community.github.io/install/project.html)? +Got everything installed? Why not start with the [tutorial project](https://leanprover-community.github.io/lean3/install/project.html)? -For more pointers, see [Learning Lean](https://leanprover-community.github.io/learn.html). +For more pointers, see [Learning Lean](https://leanprover-community.github.io/lean3/learn.html). ## Documentation Besides the installation guides above and [Lean's general -documentation](https://leanprover.github.io/documentation/), the documentation +documentation](https://leanprover.github.io/lean3/documentation/), the documentation of mathlib consists of: - [The mathlib docs](https://leanprover-community.github.io/mathlib_docs): documentation [generated @@ -33,10 +37,10 @@ of mathlib consists of: - [hole commands](https://leanprover-community.github.io/mathlib_docs/hole_commands.html), and - [attributes](https://leanprover-community.github.io/mathlib_docs/attributes.html). - A description of [currently covered theories](https://leanprover-community.github.io/theories.html), - as well as an [overview](https://leanprover-community.github.io/mathlib-overview.html) for mathematicians. + as well as an [overview](https://leanprover-community.github.io/lean3/mathlib-overview.html) for mathematicians. - A couple of [tutorial Lean files](docs/tutorial/) -- Some [extra Lean documentation](https://leanprover-community.github.io/learn.html) not specific to mathlib (see "Miscellaneous topics") -- Documentation for people who would like to [contribute to mathlib](https://leanprover-community.github.io/contribute/index.html) +- Some [extra Lean documentation](https://leanprover-community.github.io/lean3/learn.html) not specific to mathlib (see "Miscellaneous topics") +- Documentation for people who would like to [contribute to mathlib3](https://leanprover-community.github.io/lean3/contribute/index.html) Much of the discussion surrounding mathlib occurs in a [Zulip chat room](https://leanprover.zulipchat.com/). Since this @@ -49,8 +53,11 @@ welcomed. ## Contributing +> [!WARNING] +> Contributions are no longer accepted to mathlib 3; contribute to mathlib 4 instead! + The complete documentation for contributing to ``mathlib`` is located -[on the community guide contribute to mathlib](https://leanprover-community.github.io/contribute/index.html) +[on the community guide contribute to mathlib](https://leanprover-community.github.io/lean3/contribute/index.html) The process is different from other projects where one should not fork the repository. Instead write permission for non-master branches should be requested on [Zulip](https://leanprover.zulipchat.com) @@ -60,9 +67,9 @@ by introducing yourself, providing your GitHub handle and what contribution you Mathlib has the following guidelines and conventions that must be followed - - The [style guide](https://leanprover-community.github.io/contribute/style.html) - - A guide on the [naming convention](https://leanprover-community.github.io/contribute/naming.html) - - The [documentation style](https://leanprover-community.github.io/contribute/doc.html) + - The [style guide](https://leanprover-community.github.io/lean3/contribute/style.html) + - A guide on the [naming convention](https://leanprover-community.github.io/lean3/contribute/naming.html) + - The [documentation style](https://leanprover-community.github.io/lean3/contribute/doc.html) - The [commit naming conventions](https://github.com/leanprover-community/lean/blob/master/doc/commit_convention.md) Note: the title of a PR should follow the commit naming convention. From 19c869efa56bbb8b500f2724c0b77261edbfa28c Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Wed, 11 Oct 2023 09:24:47 +0200 Subject: [PATCH 1288/1291] move old README.md to OLD_README.md --- OLD_README.md | 115 ++++++++++++++++++++++++++++++++++++++++++++++++++ README.md | 110 +---------------------------------------------- 2 files changed, 116 insertions(+), 109 deletions(-) create mode 100644 OLD_README.md diff --git a/OLD_README.md b/OLD_README.md new file mode 100644 index 0000000000000..53d8c3defebf5 --- /dev/null +++ b/OLD_README.md @@ -0,0 +1,115 @@ +# Lean 3's mathlib + +> [!WARNING] +> Lean 3 and Mathlib 3 are no longer actively maintained. +> It is strongly recommended that you use [mathlib4](https://github.com/leanprover-community/mathlib4) for Lean 4 instead. + +![](https://github.com/leanprover-community/mathlib/workflows/continuous%20integration/badge.svg?branch=master) +[![Bors enabled](https://bors.tech/images/badge_small.svg)](https://app.bors.tech/repositories/24316) +[![project chat](https://img.shields.io/badge/zulip-join_chat-brightgreen.svg)](https://leanprover.zulipchat.com) +[![Gitpod Ready-to-Code](https://img.shields.io/badge/Gitpod-ready--to--code-blue?logo=gitpod)](https://gitpod.io/#https://github.com/leanprover-community/mathlib) + +[Mathlib](https://leanprover-community.github.io) is a user maintained library for the [Lean 3 theorem prover](https://github.com/leanprover-community/lean). +It contains both programming infrastructure and mathematics, +as well as tactics that use the former and allow to develop the latter. + +## Installation + +You can find detailed instructions to install Lean 3, mathlib 3, and supporting tools on [our website](https://leanprover-community.github.io/lean3/get_started.html). + +## Experimenting + +Got everything installed? Why not start with the [tutorial project](https://leanprover-community.github.io/lean3/install/project.html)? + +For more pointers, see [Learning Lean](https://leanprover-community.github.io/lean3/learn.html). + +## Documentation + +Besides the installation guides above and [Lean's general +documentation](https://leanprover.github.io/lean3/documentation/), the documentation +of mathlib consists of: + +- [The mathlib docs](https://leanprover-community.github.io/mathlib_docs): documentation [generated + automatically](https://github.com/leanprover-community/doc-gen) from the source `.lean` files. + In addition to the pages generated for each file in the library, the docs also include pages on: + - [tactics](https://leanprover-community.github.io/mathlib_docs/tactics.html), + - [commands](https://leanprover-community.github.io/mathlib_docs/commands.html), + - [hole commands](https://leanprover-community.github.io/mathlib_docs/hole_commands.html), and + - [attributes](https://leanprover-community.github.io/mathlib_docs/attributes.html). +- A description of [currently covered theories](https://leanprover-community.github.io/theories.html), + as well as an [overview](https://leanprover-community.github.io/lean3/mathlib-overview.html) for mathematicians. +- A couple of [tutorial Lean files](docs/tutorial/) +- Some [extra Lean documentation](https://leanprover-community.github.io/lean3/learn.html) not specific to mathlib (see "Miscellaneous topics") +- Documentation for people who would like to [contribute to mathlib3](https://leanprover-community.github.io/lean3/contribute/index.html) + +Much of the discussion surrounding mathlib occurs in a +[Zulip chat room](https://leanprover.zulipchat.com/). Since this +chatroom is only visible to registered users, we provide an +[openly accessible archive](https://leanprover-community.github.io/archive/) +of the public discussions. This is useful for quick reference; for a +better browsing interface, and to participate in the discussions, we strongly +suggest joining the chat. Questions from users at all levels of expertise are +welcomed. + +## Contributing + +> [!WARNING] +> Contributions are no longer accepted to mathlib 3; contribute to mathlib 4 instead! + +The complete documentation for contributing to ``mathlib`` is located +[on the community guide contribute to mathlib](https://leanprover-community.github.io/lean3/contribute/index.html) + +The process is different from other projects where one should not fork the repository. +Instead write permission for non-master branches should be requested on [Zulip](https://leanprover.zulipchat.com) +by introducing yourself, providing your GitHub handle and what contribution you are planning on doing. + +### Guidelines + +Mathlib has the following guidelines and conventions that must be followed + + - The [style guide](https://leanprover-community.github.io/lean3/contribute/style.html) + - A guide on the [naming convention](https://leanprover-community.github.io/lean3/contribute/naming.html) + - The [documentation style](https://leanprover-community.github.io/lean3/contribute/doc.html) + - The [commit naming conventions](https://github.com/leanprover-community/lean/blob/master/doc/commit_convention.md) + +Note: the title of a PR should follow the commit naming convention. + +### Using ``leanproject`` to contribute + +Running the ``leanproject get -b mathlib:shiny_lemma`` command will create a new worktree ``mathlib_shiny_lemma`` +with a local branch called ``shiny_lemma`` which has a copy of mathlib to work on. + +``leanproject build`` will check that nothing broke. +Be warned that this will take some time if a fundamental file was changed. + +## Maintainers: + +* Anne Baanen (@Vierkantor): algebra, number theory, tactics +* Reid Barton (@rwbarton): category theory, topology +* Riccardo Brasca (@riccardobrasca): algebra, number theory, algebraic geometry, category theory +* Mario Carneiro (@digama0): lean formalization, tactics, type theory, proof engineering +* Bryan Gin-ge Chen (@bryangingechen): documentation, infrastructure +* Johan Commelin (@jcommelin): algebra, number theory, category theory, algebraic geometry +* Rémy Degenne (@RemyDegenne): probability, measure theory, analysis +* Floris van Doorn (@fpvandoorn): measure theory, model theory, tactics +* Frédéric Dupuis (@dupuisf): linear algebra, functional analysis +* Gabriel Ebner (@gebner): tactics, infrastructure, core, formal languages +* Sébastien Gouëzel (@sgouezel): topology, calculus, geometry, analysis, measure theory +* Markus Himmel (@TwoFX): category theory +* Chris Hughes (@ChrisHughes24): algebra +* Yury G. Kudryashov (@urkud): analysis, topology, measure theory +* Robert Y. Lewis (@robertylewis): tactics, documentation +* Heather Macbeth (@hrmacbeth): geometry, analysis +* Patrick Massot (@patrickmassot): documentation, topology, geometry +* Bhavik Mehta (@b-mehta): category theory, combinatorics +* Kyle Miller (@kmill): combinatorics, documentation +* Scott Morrison (@semorrison): category theory, tactics +* Oliver Nash (@ocfnash): algebra, geometry, topology +* Adam Topaz (@adamtopaz): algebra, category theory, algebraic geometry +* Eric Wieser (@eric-wieser): algebra, infrastructure + +## Emeritus maintainers: + +* Jeremy Avigad (@avigad): analysis +* Johannes Hölzl (@johoelzl): measure theory, topology +* Simon Hudon (@cipher1024): tactics diff --git a/README.md b/README.md index 53d8c3defebf5..f850c525cc3b0 100644 --- a/README.md +++ b/README.md @@ -4,112 +4,4 @@ > Lean 3 and Mathlib 3 are no longer actively maintained. > It is strongly recommended that you use [mathlib4](https://github.com/leanprover-community/mathlib4) for Lean 4 instead. -![](https://github.com/leanprover-community/mathlib/workflows/continuous%20integration/badge.svg?branch=master) -[![Bors enabled](https://bors.tech/images/badge_small.svg)](https://app.bors.tech/repositories/24316) -[![project chat](https://img.shields.io/badge/zulip-join_chat-brightgreen.svg)](https://leanprover.zulipchat.com) -[![Gitpod Ready-to-Code](https://img.shields.io/badge/Gitpod-ready--to--code-blue?logo=gitpod)](https://gitpod.io/#https://github.com/leanprover-community/mathlib) - -[Mathlib](https://leanprover-community.github.io) is a user maintained library for the [Lean 3 theorem prover](https://github.com/leanprover-community/lean). -It contains both programming infrastructure and mathematics, -as well as tactics that use the former and allow to develop the latter. - -## Installation - -You can find detailed instructions to install Lean 3, mathlib 3, and supporting tools on [our website](https://leanprover-community.github.io/lean3/get_started.html). - -## Experimenting - -Got everything installed? Why not start with the [tutorial project](https://leanprover-community.github.io/lean3/install/project.html)? - -For more pointers, see [Learning Lean](https://leanprover-community.github.io/lean3/learn.html). - -## Documentation - -Besides the installation guides above and [Lean's general -documentation](https://leanprover.github.io/lean3/documentation/), the documentation -of mathlib consists of: - -- [The mathlib docs](https://leanprover-community.github.io/mathlib_docs): documentation [generated - automatically](https://github.com/leanprover-community/doc-gen) from the source `.lean` files. - In addition to the pages generated for each file in the library, the docs also include pages on: - - [tactics](https://leanprover-community.github.io/mathlib_docs/tactics.html), - - [commands](https://leanprover-community.github.io/mathlib_docs/commands.html), - - [hole commands](https://leanprover-community.github.io/mathlib_docs/hole_commands.html), and - - [attributes](https://leanprover-community.github.io/mathlib_docs/attributes.html). -- A description of [currently covered theories](https://leanprover-community.github.io/theories.html), - as well as an [overview](https://leanprover-community.github.io/lean3/mathlib-overview.html) for mathematicians. -- A couple of [tutorial Lean files](docs/tutorial/) -- Some [extra Lean documentation](https://leanprover-community.github.io/lean3/learn.html) not specific to mathlib (see "Miscellaneous topics") -- Documentation for people who would like to [contribute to mathlib3](https://leanprover-community.github.io/lean3/contribute/index.html) - -Much of the discussion surrounding mathlib occurs in a -[Zulip chat room](https://leanprover.zulipchat.com/). Since this -chatroom is only visible to registered users, we provide an -[openly accessible archive](https://leanprover-community.github.io/archive/) -of the public discussions. This is useful for quick reference; for a -better browsing interface, and to participate in the discussions, we strongly -suggest joining the chat. Questions from users at all levels of expertise are -welcomed. - -## Contributing - -> [!WARNING] -> Contributions are no longer accepted to mathlib 3; contribute to mathlib 4 instead! - -The complete documentation for contributing to ``mathlib`` is located -[on the community guide contribute to mathlib](https://leanprover-community.github.io/lean3/contribute/index.html) - -The process is different from other projects where one should not fork the repository. -Instead write permission for non-master branches should be requested on [Zulip](https://leanprover.zulipchat.com) -by introducing yourself, providing your GitHub handle and what contribution you are planning on doing. - -### Guidelines - -Mathlib has the following guidelines and conventions that must be followed - - - The [style guide](https://leanprover-community.github.io/lean3/contribute/style.html) - - A guide on the [naming convention](https://leanprover-community.github.io/lean3/contribute/naming.html) - - The [documentation style](https://leanprover-community.github.io/lean3/contribute/doc.html) - - The [commit naming conventions](https://github.com/leanprover-community/lean/blob/master/doc/commit_convention.md) - -Note: the title of a PR should follow the commit naming convention. - -### Using ``leanproject`` to contribute - -Running the ``leanproject get -b mathlib:shiny_lemma`` command will create a new worktree ``mathlib_shiny_lemma`` -with a local branch called ``shiny_lemma`` which has a copy of mathlib to work on. - -``leanproject build`` will check that nothing broke. -Be warned that this will take some time if a fundamental file was changed. - -## Maintainers: - -* Anne Baanen (@Vierkantor): algebra, number theory, tactics -* Reid Barton (@rwbarton): category theory, topology -* Riccardo Brasca (@riccardobrasca): algebra, number theory, algebraic geometry, category theory -* Mario Carneiro (@digama0): lean formalization, tactics, type theory, proof engineering -* Bryan Gin-ge Chen (@bryangingechen): documentation, infrastructure -* Johan Commelin (@jcommelin): algebra, number theory, category theory, algebraic geometry -* Rémy Degenne (@RemyDegenne): probability, measure theory, analysis -* Floris van Doorn (@fpvandoorn): measure theory, model theory, tactics -* Frédéric Dupuis (@dupuisf): linear algebra, functional analysis -* Gabriel Ebner (@gebner): tactics, infrastructure, core, formal languages -* Sébastien Gouëzel (@sgouezel): topology, calculus, geometry, analysis, measure theory -* Markus Himmel (@TwoFX): category theory -* Chris Hughes (@ChrisHughes24): algebra -* Yury G. Kudryashov (@urkud): analysis, topology, measure theory -* Robert Y. Lewis (@robertylewis): tactics, documentation -* Heather Macbeth (@hrmacbeth): geometry, analysis -* Patrick Massot (@patrickmassot): documentation, topology, geometry -* Bhavik Mehta (@b-mehta): category theory, combinatorics -* Kyle Miller (@kmill): combinatorics, documentation -* Scott Morrison (@semorrison): category theory, tactics -* Oliver Nash (@ocfnash): algebra, geometry, topology -* Adam Topaz (@adamtopaz): algebra, category theory, algebraic geometry -* Eric Wieser (@eric-wieser): algebra, infrastructure - -## Emeritus maintainers: - -* Jeremy Avigad (@avigad): analysis -* Johannes Hölzl (@johoelzl): measure theory, topology -* Simon Hudon (@cipher1024): tactics +(If you need to read the old `README.md`, please see `OLD_README.md`.) From b1abe23ae96fef89ad30d9f4362c307f72a55010 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 16 Oct 2023 09:24:06 +0000 Subject: [PATCH 1289/1291] =?UTF-8?q?feat(measure=5Ftheory/order/upper=5Fl?= =?UTF-8?q?ower):=20Order-connected=20sets=20in=20`=E2=84=9D=E2=81=BF`=20a?= =?UTF-8?q?re=20measurable=20(#16976)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Prove that the frontier of an order-connected set in `ℝⁿ` (with the `∞`-metric, but it doesn't actually matter) has measure zero. As a corollary, antichains in `ℝⁿ` have measure zero. Co-authored-by: @JasonKYi --- src/analysis/normed/order/upper_lower.lean | 161 +++++++++++++++++++- src/data/real/nnreal.lean | 2 + src/measure_theory/order/upper_lower.lean | 131 ++++++++++++++++ src/order/bounds/basic.lean | 2 + src/topology/algebra/order/upper_lower.lean | 29 +++- 5 files changed, 320 insertions(+), 5 deletions(-) create mode 100644 src/measure_theory/order/upper_lower.lean diff --git a/src/analysis/normed/order/upper_lower.lean b/src/analysis/normed/order/upper_lower.lean index c834d2d671ddc..a081b70aa9623 100644 --- a/src/analysis/normed/order/upper_lower.lean +++ b/src/analysis/normed/order/upper_lower.lean @@ -23,10 +23,11 @@ are measurable. -/ open function metric set +open_locale pointwise variables {α ι : Type*} -section metric_space +section normed_ordered_group variables [normed_ordered_group α] {s : set α} @[to_additive is_upper_set.thickening] @@ -49,12 +50,22 @@ protected lemma is_lower_set.cthickening' (hs : is_lower_set s) (ε : ℝ) : is_lower_set (cthickening ε s) := by { rw cthickening_eq_Inter_thickening'', exact is_lower_set_Inter₂ (λ δ hδ, hs.thickening' _) } -end metric_space +@[to_additive upper_closure_interior_subset] +lemma upper_closure_interior_subset' (s : set α) : + (upper_closure (interior s) : set α) ⊆ interior (upper_closure s) := +upper_closure_min (interior_mono subset_upper_closure) (upper_closure s).upper.interior + +@[to_additive lower_closure_interior_subset] +lemma lower_closure_interior_subset' (s : set α) : + (upper_closure (interior s) : set α) ⊆ interior (upper_closure s) := +upper_closure_min (interior_mono subset_upper_closure) (upper_closure s).upper.interior + +end normed_ordered_group /-! ### `ℝⁿ` -/ section finite -variables [finite ι] {s : set (ι → ℝ)} {x y : ι → ℝ} {δ : ℝ} +variables [finite ι] {s : set (ι → ℝ)} {x y : ι → ℝ} lemma is_upper_set.mem_interior_of_forall_lt (hs : is_upper_set s) (hx : x ∈ closure s) (h : ∀ i, x i < y i) : @@ -99,7 +110,78 @@ end end finite section fintype -variables [fintype ι] {s : set (ι → ℝ)} {x y : ι → ℝ} {δ : ℝ} +variables [fintype ι] {s t : set (ι → ℝ)} {a₁ a₂ b₁ b₂ x y : ι → ℝ} {δ : ℝ} + +-- TODO: Generalise those lemmas so that they also apply to `ℝ` and `euclidean_space ι ℝ` +lemma dist_inf_sup (x y : ι → ℝ) : dist (x ⊓ y) (x ⊔ y) = dist x y := +begin + refine congr_arg coe (finset.sup_congr rfl $ λ i _, _), + simp only [real.nndist_eq', sup_eq_max, inf_eq_min, max_sub_min_eq_abs, pi.inf_apply, + pi.sup_apply, real.nnabs_of_nonneg, abs_nonneg, real.to_nnreal_abs], +end + +lemma dist_mono_left : monotone_on (λ x, dist x y) (Ici y) := +begin + refine λ y₁ hy₁ y₂ hy₂ hy, nnreal.coe_le_coe.2 (finset.sup_mono_fun $ λ i _, _), + rw [real.nndist_eq, real.nnabs_of_nonneg (sub_nonneg_of_le (‹y ≤ _› i : y i ≤ y₁ i)), + real.nndist_eq, real.nnabs_of_nonneg (sub_nonneg_of_le (‹y ≤ _› i : y i ≤ y₂ i))], + exact real.to_nnreal_mono (sub_le_sub_right (hy _) _), +end + +lemma dist_mono_right : monotone_on (dist x) (Ici x) := +by simpa only [dist_comm] using dist_mono_left + +lemma dist_anti_left : antitone_on (λ x, dist x y) (Iic y) := +begin + refine λ y₁ hy₁ y₂ hy₂ hy, nnreal.coe_le_coe.2 (finset.sup_mono_fun $ λ i _, _), + rw [real.nndist_eq', real.nnabs_of_nonneg (sub_nonneg_of_le (‹_ ≤ y› i : y₂ i ≤ y i)), + real.nndist_eq', real.nnabs_of_nonneg (sub_nonneg_of_le (‹_ ≤ y› i : y₁ i ≤ y i))], + exact real.to_nnreal_mono (sub_le_sub_left (hy _) _), +end + +lemma dist_anti_right : antitone_on (dist x) (Iic x) := +by simpa only [dist_comm] using dist_anti_left + +lemma dist_le_dist_of_le (ha : a₂ ≤ a₁) (h₁ : a₁ ≤ b₁) (hb : b₁ ≤ b₂) : dist a₁ b₁ ≤ dist a₂ b₂ := +(dist_mono_right h₁ (h₁.trans hb) hb).trans $ + dist_anti_left (ha.trans $ h₁.trans hb) (h₁.trans hb) ha + +protected lemma metric.bounded.bdd_below : bounded s → bdd_below s := +begin + rintro ⟨r, hr⟩, + obtain rfl | ⟨x, hx⟩ := s.eq_empty_or_nonempty, + { exact bdd_below_empty }, + { exact ⟨x - const _ r, λ y hy i, sub_le_comm.1 + (abs_sub_le_iff.1 $ (dist_le_pi_dist _ _ _).trans $ hr _ hx _ hy).1⟩ } +end + +protected lemma metric.bounded.bdd_above : bounded s → bdd_above s := +begin + rintro ⟨r, hr⟩, + obtain rfl | ⟨x, hx⟩ := s.eq_empty_or_nonempty, + { exact bdd_above_empty }, + { exact ⟨x + const _ r, λ y hy i, sub_le_iff_le_add'.1 $ + (abs_sub_le_iff.1 $ (dist_le_pi_dist _ _ _).trans $ hr _ hx _ hy).2⟩ } +end + +protected lemma bdd_below.bounded : bdd_below s → bdd_above s → bounded s := +begin + rintro ⟨a, ha⟩ ⟨b, hb⟩, + refine ⟨dist a b, λ x hx y hy, _⟩, + rw ←dist_inf_sup, + exact dist_le_dist_of_le (le_inf (ha hx) $ ha hy) inf_le_sup (sup_le (hb hx) $ hb hy), +end + +protected lemma bdd_above.bounded : bdd_above s → bdd_below s → bounded s := flip bdd_below.bounded + +lemma bounded_iff_bdd_below_bdd_above : bounded s ↔ bdd_below s ∧ bdd_above s := +⟨λ h, ⟨h.bdd_below, h.bdd_above⟩, λ h, h.1.bounded h.2⟩ + +lemma bdd_below.bounded_inter (hs : bdd_below s) (ht : bdd_above t) : bounded (s ∩ t) := +(hs.mono $ inter_subset_left _ _).bounded $ ht.mono $ inter_subset_right _ _ + +lemma bdd_above.bounded_inter (hs : bdd_above s) (ht : bdd_below t) : bounded (s ∩ t) := +(hs.mono $ inter_subset_left _ _).bounded $ ht.mono $ inter_subset_right _ _ lemma is_upper_set.exists_subset_ball (hs : is_upper_set s) (hx : x ∈ closure s) (hδ : 0 < δ) : ∃ y, closed_ball y (δ/4) ⊆ closed_ball x δ ∧ closed_ball y (δ/4) ⊆ interior s := @@ -140,3 +222,74 @@ begin end end fintype + +section finite +variables [finite ι] {s t : set (ι → ℝ)} {a₁ a₂ b₁ b₂ x y : ι → ℝ} {δ : ℝ} + +lemma is_antichain.interior_eq_empty [nonempty ι] (hs : is_antichain (≤) s) : interior s = ∅ := +begin + casesI nonempty_fintype ι, + refine eq_empty_of_forall_not_mem (λ x hx, _), + have hx' := interior_subset hx, + rw [mem_interior_iff_mem_nhds, metric.mem_nhds_iff] at hx, + obtain ⟨ε, hε, hx⟩ := hx, + refine hs.not_lt hx' (hx _) (lt_add_of_pos_right _ (by positivity : 0 < const ι (ε / 2))), + simpa [const, @pi_norm_const ι ℝ _ _ _ (ε / 2), abs_of_nonneg hε.lt.le], +end + +/-! +#### Note + +The closure and frontier of an antichain might not be antichains. Take for example the union +of the open segments from `(0, 2)` to `(1, 1)` and from `(2, 1)` to `(3, 0)`. `(1, 1)` and `(2, 1)` +are comparable and both in the closure/frontier. +-/ + +protected lemma is_closed.upper_closure (hs : is_closed s) (hs' : bdd_below s) : + is_closed (upper_closure s : set (ι → ℝ)) := +begin + casesI nonempty_fintype ι, + refine is_seq_closed.is_closed (λ f x hf hx, _), + choose g hg hgf using hf, + obtain ⟨a, ha⟩ := hx.bdd_above_range, + obtain ⟨b, hb, φ, hφ, hbf⟩ := tendsto_subseq_of_bounded (hs'.bounded_inter + bdd_above_Iic) (λ n, ⟨hg n, (hgf _).trans $ ha $ mem_range_self _⟩), + exact ⟨b, closure_minimal (inter_subset_left _ _) hs hb, + le_of_tendsto_of_tendsto' hbf (hx.comp hφ.tendsto_at_top) $ λ _, hgf _⟩, +end + +protected lemma is_closed.lower_closure (hs : is_closed s) (hs' : bdd_above s) : + is_closed (lower_closure s : set (ι → ℝ)) := +begin + casesI nonempty_fintype ι, + refine is_seq_closed.is_closed (λ f x hf hx, _), + choose g hg hfg using hf, + haveI : bounded_ge_nhds_class ℝ := by apply_instance, + obtain ⟨a, ha⟩ := hx.bdd_below_range, + obtain ⟨b, hb, φ, hφ, hbf⟩ := tendsto_subseq_of_bounded (hs'.bounded_inter + bdd_below_Ici) (λ n, ⟨hg n, (ha $ mem_range_self _).trans $ hfg _⟩), + exact ⟨b, closure_minimal (inter_subset_left _ _) hs hb, + le_of_tendsto_of_tendsto' (hx.comp hφ.tendsto_at_top) hbf $ λ _, hfg _⟩, +end + +protected lemma is_clopen.upper_closure (hs : is_clopen s) (hs' : bdd_below s) : + is_clopen (upper_closure s : set (ι → ℝ)) := +⟨hs.1.upper_closure, hs.2.upper_closure hs'⟩ + +protected lemma is_clopen.lower_closure (hs : is_clopen s) (hs' : bdd_above s) : + is_clopen (lower_closure s : set (ι → ℝ)) := +⟨hs.1.lower_closure, hs.2.lower_closure hs'⟩ + +lemma closure_upper_closure_comm (hs : bdd_below s) : + closure (upper_closure s : set (ι → ℝ)) = upper_closure (closure s) := +(closure_minimal (upper_closure_anti subset_closure) $ + is_closed_closure.upper_closure hs.closure).antisymm $ + upper_closure_min (closure_mono subset_upper_closure) (upper_closure s).upper.closure + +lemma closure_lower_closure_comm (hs : bdd_above s) : + closure (lower_closure s : set (ι → ℝ)) = lower_closure (closure s) := +(closure_minimal (lower_closure_mono subset_closure) $ + is_closed_closure.lower_closure hs.closure).antisymm $ + lower_closure_min (closure_mono subset_lower_closure) (lower_closure s).lower.closure + +end finite diff --git a/src/data/real/nnreal.lean b/src/data/real/nnreal.lean index cca419cf48609..47dba7306a2d2 100644 --- a/src/data/real/nnreal.lean +++ b/src/data/real/nnreal.lean @@ -822,6 +822,8 @@ lemma nnabs_coe (x : ℝ≥0) : nnabs x = x := by simp lemma coe_to_nnreal_le (x : ℝ) : (to_nnreal x : ℝ) ≤ |x| := max_le (le_abs_self _) (abs_nonneg _) +@[simp] lemma to_nnreal_abs (x : ℝ) : |x|.to_nnreal = x.nnabs := nnreal.coe_injective $ by simp + lemma cast_nat_abs_eq_nnabs_cast (n : ℤ) : (n.nat_abs : ℝ≥0) = nnabs n := by { ext, rw [nnreal.coe_nat_cast, int.cast_nat_abs, real.coe_nnabs] } diff --git a/src/measure_theory/order/upper_lower.lean b/src/measure_theory/order/upper_lower.lean new file mode 100644 index 0000000000000..d6bc49a4cfffa --- /dev/null +++ b/src/measure_theory/order/upper_lower.lean @@ -0,0 +1,131 @@ +/- +Copyright (c) 2022 Yaël Dillies, Kexing Ying. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies, Kexing Ying +-/ +import analysis.normed.order.upper_lower +import logic.lemmas +import measure_theory.covering.besicovitch_vector_space + +/-! +# Order-connected sets are null-measurable + +This file proves that order-connected sets in `ℝⁿ` under the pointwise order are null-measurable. +Recall that `x ≤ y` iff `∀ i, x i ≤ y i`, and `s` is order-connected iff +`∀ x y ∈ s, ∀ z, x ≤ z → z ≤ y → z ∈ s`. + +## Main declarations + +* `set.ord_connected.null_frontier`: The frontier of an order-connected set in `ℝⁿ` has measure `0`. + +## Notes + +We prove null-measurability in `ℝⁿ` with the `∞`-metric, but this transfers directly to `ℝⁿ` with +the Euclidean metric because they have the same measurable sets. + +Null-measurability can't be strengthened to measurability because any antichain (and in particular +any subset of the antidiagonal `{(x, y) | x + y = 0}`) is order-connected. + +## TODO + +Generalize so that it also applies to `ℝ × ℝ`, for example. +-/ + +open filter measure_theory metric set +open_locale topology + +variables {ι : Type*} [fintype ι] {s : set (ι → ℝ)} {x y : ι → ℝ} {δ : ℝ} + +/-- If we can fit a small ball inside a set `s` intersected with any neighborhood of `x`, then the +density of `s` near `x` is not `0`. Along with `aux₁`, this proves that `x` is a Lebesgue point of +`s`. This will be used to prove that the frontier of an order-connected set is null. -/ +private lemma aux₀ + (h : ∀ δ, 0 < δ → ∃ y, closed_ball y (δ/4) ⊆ closed_ball x δ ∧ closed_ball y (δ/4) ⊆ interior s) : + ¬ tendsto (λ r, volume (closure s ∩ closed_ball x r) / volume (closed_ball x r)) (𝓝[>] 0) + (𝓝 0) := +begin + choose f hf₀ hf₁ using h, + intros H, + obtain ⟨ε, hε, hε', hε₀⟩ := exists_seq_strict_anti_tendsto_nhds_within (0 : ℝ), + refine not_eventually.2 (frequently_of_forall $ λ _, lt_irrefl $ + ennreal.of_real $ 4⁻¹ ^ fintype.card ι) + ((tendsto.eventually_lt (H.comp hε₀) tendsto_const_nhds _).mono $ λ n, lt_of_le_of_lt _), + swap, + refine (ennreal.div_le_div_right (volume.mono $ subset_inter + ((hf₁ _ $ hε' n).trans interior_subset_closure) $ hf₀ _ $ hε' n) _).trans_eq' _, + dsimp, + have := hε' n, + rw [real.volume_pi_closed_ball, real.volume_pi_closed_ball, ←ennreal.of_real_div_of_pos, ←div_pow, + mul_div_mul_left _ _ (two_ne_zero' ℝ), div_right_comm, div_self, one_div], + all_goals { positivity }, +end + +/-- If we can fit a small ball inside a set `sᶜ` intersected with any neighborhood of `x`, then the +density of `s` near `x` is not `1`. Along with `aux₀`, this proves that `x` is a Lebesgue point of +`s`. This will be used to prove that the frontier of an order-connected set is null. -/ +private lemma aux₁ + (h : ∀ δ, 0 < δ → + ∃ y, closed_ball y (δ/4) ⊆ closed_ball x δ ∧ closed_ball y (δ/4) ⊆ interior sᶜ) : + ¬ tendsto (λ r, volume (closure s ∩ closed_ball x r) / volume (closed_ball x r)) (𝓝[>] 0) + (𝓝 1) := +begin + choose f hf₀ hf₁ using h, + intros H, + obtain ⟨ε, hε, hε', hε₀⟩ := exists_seq_strict_anti_tendsto_nhds_within (0 : ℝ), + refine not_eventually.2 (frequently_of_forall $ λ _, lt_irrefl $ + 1 - ennreal.of_real (4⁻¹ ^ fintype.card ι)) + ((tendsto.eventually_lt tendsto_const_nhds (H.comp hε₀) $ + ennreal.sub_lt_self ennreal.one_ne_top one_ne_zero _).mono $ λ n, lt_of_le_of_lt' _), + swap, + refine (ennreal.div_le_div_right (volume.mono _) _).trans_eq _, + { exact closed_ball x (ε n) \ closed_ball (f (ε n) $ hε' n) (ε n / 4) }, + { rw diff_eq_compl_inter, + refine inter_subset_inter_left _ _, + rw [subset_compl_comm, ←interior_compl], + exact hf₁ _ _ }, + dsimp, + have := hε' n, + rw [measure_diff (hf₀ _ _) _ ((real.volume_pi_closed_ball _ _).trans_ne ennreal.of_real_ne_top), + real.volume_pi_closed_ball, real.volume_pi_closed_ball, ennreal.sub_div (λ _ _, _), + ennreal.div_self _ ennreal.of_real_ne_top, ←ennreal.of_real_div_of_pos, ←div_pow, + mul_div_mul_left _ _ (two_ne_zero' ℝ), div_right_comm, div_self, one_div], + all_goals { positivity <|> measurability }, +end + +lemma is_upper_set.null_frontier (hs : is_upper_set s) : volume (frontier s) = 0 := +begin + refine eq_bot_mono (volume.mono $ λ x hx, _) + (besicovitch.ae_tendsto_measure_inter_div_of_measurable_set _ is_closed_closure.measurable_set), + { exact s }, + by_cases x ∈ closure s; simp [h], + { exact aux₁ (λ _, hs.compl.exists_subset_ball $ frontier_subset_closure $ + by rwa frontier_compl) }, + { exact aux₀ (λ _, hs.exists_subset_ball $ frontier_subset_closure hx) } +end + +lemma is_lower_set.null_frontier (hs : is_lower_set s) : volume (frontier s) = 0 := +begin + refine eq_bot_mono (volume.mono $ λ x hx, _) + (besicovitch.ae_tendsto_measure_inter_div_of_measurable_set _ is_closed_closure.measurable_set), + { exact s }, + by_cases x ∈ closure s; simp [h], + { exact aux₁ (λ _, hs.compl.exists_subset_ball $ frontier_subset_closure $ + by rwa frontier_compl) }, + { exact aux₀ (λ _, hs.exists_subset_ball $ frontier_subset_closure hx) } +end + +lemma set.ord_connected.null_frontier (hs : s.ord_connected) : volume (frontier s) = 0 := +begin + rw ← hs.upper_closure_inter_lower_closure, + refine le_bot_iff.1 ((volume.mono $ (frontier_inter_subset _ _).trans $ union_subset_union + (inter_subset_left _ _) $ inter_subset_right _ _).trans $ (measure_union_le _ _).trans_eq _), + rw [(upper_set.upper _).null_frontier, (lower_set.lower _).null_frontier, zero_add, bot_eq_zero], +end + +protected lemma set.ord_connected.null_measurable_set (hs : s.ord_connected) : + null_measurable_set s := +null_measurable_set_of_null_frontier hs.null_frontier + +lemma is_antichain.volume_eq_zero [nonempty ι] (hs : is_antichain (≤) s) : volume s = 0 := +le_bot_iff.1 $ (volume.mono $ by { rw [←closure_diff_interior, hs.interior_eq_empty, diff_empty], + exact subset_closure }).trans_eq hs.ord_connected.null_frontier diff --git a/src/order/bounds/basic.lean b/src/order/bounds/basic.lean index 6add0bb14d555..ab986bd7f2fd4 100644 --- a/src/order/bounds/basic.lean +++ b/src/order/bounds/basic.lean @@ -62,6 +62,8 @@ def is_glb (s : set α) : α → Prop := is_greatest (lower_bounds s) lemma mem_upper_bounds : a ∈ upper_bounds s ↔ ∀ x ∈ s, x ≤ a := iff.rfl lemma mem_lower_bounds : a ∈ lower_bounds s ↔ ∀ x ∈ s, a ≤ x := iff.rfl +lemma mem_upper_bounds_iff_subset_Iic : a ∈ upper_bounds s ↔ s ⊆ Iic a := iff.rfl +lemma mem_lower_bounds_iff_subset_Ici : a ∈ lower_bounds s ↔ s ⊆ Ici a := iff.rfl lemma bdd_above_def : bdd_above s ↔ ∃ x, ∀ y ∈ s, y ≤ x := iff.rfl lemma bdd_below_def : bdd_below s ↔ ∃ x, ∀ y ∈ s, x ≤ y := iff.rfl diff --git a/src/topology/algebra/order/upper_lower.lean b/src/topology/algebra/order/upper_lower.lean index 792ad0b431a15..c1286acac2cdc 100644 --- a/src/topology/algebra/order/upper_lower.lean +++ b/src/topology/algebra/order/upper_lower.lean @@ -5,6 +5,7 @@ Authors: Yaël Dillies -/ import algebra.order.upper_lower import topology.algebra.group.basic +import topology.order.basic /-! # Topological facts about upper/lower/order-connected sets @@ -46,7 +47,33 @@ instance ordered_comm_group.to_has_upper_lower_closure [ordered_comm_group α] is_open_upper_closure := λ s hs, by { rw [←mul_one s, ←mul_upper_closure], exact hs.mul_right }, is_open_lower_closure := λ s hs, by { rw [←mul_one s, ←mul_lower_closure], exact hs.mul_right } } -variables [preorder α] [has_upper_lower_closure α] {s : set α} +variables [preorder α] + +section order_closed_topology +variables [order_closed_topology α] {s : set α} + +@[simp] lemma upper_bounds_closure (s : set α) : + upper_bounds (closure s : set α) = upper_bounds s := +ext $ λ a, by simp_rw [mem_upper_bounds_iff_subset_Iic, is_closed_Iic.closure_subset_iff] + +@[simp] lemma lower_bounds_closure (s : set α) : + lower_bounds (closure s : set α) = lower_bounds s := +ext $ λ a, by simp_rw [mem_lower_bounds_iff_subset_Ici, is_closed_Ici.closure_subset_iff] + +@[simp] lemma bdd_above_closure : bdd_above (closure s) ↔ bdd_above s := +by simp_rw [bdd_above, upper_bounds_closure] + +@[simp] lemma bdd_below_closure : bdd_below (closure s) ↔ bdd_below s := +by simp_rw [bdd_below, lower_bounds_closure] + +alias bdd_above_closure ↔ bdd_above.of_closure bdd_above.closure +alias bdd_below_closure ↔ bdd_below.of_closure bdd_below.closure + +attribute [protected] bdd_above.closure bdd_below.closure + +end order_closed_topology + +variables [has_upper_lower_closure α] {s : set α} protected lemma is_upper_set.closure : is_upper_set s → is_upper_set (closure s) := has_upper_lower_closure.is_upper_set_closure _ From 3365b20c2ffa7c35e47e5209b89ba9abdddf3ffe Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sun, 22 Oct 2023 19:51:38 +0000 Subject: [PATCH 1290/1291] feat(combinatorics/simple_graph): More clique lemmas (#19203) More lemmas about `is_clique`, `is_n_clique`, `edge_set`. Also define `clique_free_on`, a local version of `clique_free`. --- src/combinatorics/simple_graph/basic.lean | 122 +++++++++- src/combinatorics/simple_graph/clique.lean | 211 ++++++++++++++++-- .../simple_graph/triangle/basic.lean | 3 +- src/data/finset/card.lean | 7 + src/data/finset/preimage.lean | 4 + src/logic/basic.lean | 9 + src/logic/relation.lean | 23 +- 7 files changed, 349 insertions(+), 30 deletions(-) diff --git a/src/combinatorics/simple_graph/basic.lean b/src/combinatorics/simple_graph/basic.lean index 5c66632b5e4b5..0b1d706d9e560 100644 --- a/src/combinatorics/simple_graph/basic.lean +++ b/src/combinatorics/simple_graph/basic.lean @@ -3,6 +3,7 @@ Copyright (c) 2020 Aaron Anderson, Jalex Stark, Kyle Miller. All rights reserved Released under Apache 2.0 license as described in the file LICENSE. Authors: Aaron Anderson, Jalex Stark, Kyle Miller, Alena Gusakov, Hunter Monroe -/ +import data.fun_like.fintype import data.rel import data.set.finite import data.sym.sym2 @@ -288,6 +289,13 @@ instance : complete_boolean_algebra (simple_graph V) := @[simps] instance (V : Type u) : inhabited (simple_graph V) := ⟨⊥⟩ +instance [subsingleton V] : unique (simple_graph V) := +{ default := ⊥, + uniq := λ G, by ext a b; simp [subsingleton.elim a b] } + +instance [nontrivial V] : nontrivial (simple_graph V) := +⟨⟨⊥, ⊤, λ h, not_subsingleton V ⟨by simpa [ext_iff, function.funext_iff] using h⟩⟩⟩ + section decidable variables (V) (H : simple_graph V) [decidable_rel G.adj] [decidable_rel H.adj] @@ -372,6 +380,17 @@ by { ext ⟨x, y⟩, refl } @[simp] lemma edge_set_sdiff : (G₁ \ G₂).edge_set = G₁.edge_set \ G₂.edge_set := by { ext ⟨x, y⟩, refl } +variables {G G₁ G₂} + +@[simp] lemma disjoint_edge_set : disjoint G₁.edge_set G₂.edge_set ↔ disjoint G₁ G₂ := +by rw [set.disjoint_iff, disjoint_iff_inf_le, ←edge_set_inf, ←edge_set_bot, ←set.le_iff_subset, + order_embedding.le_iff_le] + +@[simp] lemma edge_set_eq_empty : G.edge_set = ∅ ↔ G = ⊥ := by rwa [←edge_set_bot, edge_set_inj] + +@[simp] lemma edge_set_nonempty : G.edge_set.nonempty ↔ G ≠ ⊥ := +by rw [set.nonempty_iff_ne_empty, edge_set_eq_empty.ne] + /-- This lemma, combined with `edge_set_sdiff` and `edge_set_from_edge_set`, allows proving `(G \ from_edge_set s).edge_set = G.edge_set \ s` by `simp`. @@ -406,6 +425,8 @@ end lemma adj_iff_exists_edge_coe : G.adj a b ↔ ∃ (e : G.edge_set), ↑e = ⟦(a, b)⟧ := by simp only [mem_edge_set, exists_prop, set_coe.exists, exists_eq_right, subtype.coe_mk] +variables (G G₁ G₂) + lemma edge_other_ne {e : sym2 V} (he : e ∈ G.edge_set) {v : V} (h : v ∈ e) : h.other ≠ v := begin erw [← sym2.other_spec h, sym2.eq_swap] at he, @@ -482,6 +503,16 @@ begin exact λ vws _, h vws, end +@[simp] lemma disjoint_from_edge_set : disjoint G (from_edge_set s) ↔ disjoint G.edge_set s := +begin + conv_rhs { rw ←set.diff_union_inter s {e | e.is_diag} }, + rw [←disjoint_edge_set, edge_set_from_edge_set, set.disjoint_union_right, and_iff_left], + exact set.disjoint_left.2 (λ e he he', not_is_diag_of_mem_edge_set _ he he'.2), +end + +@[simp] lemma from_edge_set_disjoint : disjoint (from_edge_set s) G ↔ disjoint s G.edge_set := +by rw [disjoint.comm, disjoint_from_edge_set, disjoint.comm] + instance [decidable_eq V] [fintype s] : fintype (from_edge_set s).edge_set := by { rw edge_set_from_edge_set s, apply_instance } @@ -680,6 +711,9 @@ end edge_finset @[simp] lemma mem_neighbor_set (v w : V) : w ∈ G.neighbor_set v ↔ G.adj v w := iff.rfl +@[simp] lemma not_mem_neighbor_set_self : a ∉ G.neighbor_set a := +(mem_neighbor_set _ _ _).not.2 $ G.loopless _ + @[simp] lemma mem_incidence_set (v w : V) : ⟦(v, w)⟧ ∈ G.incidence_set v ↔ G.adj v w := by simp [incidence_set] @@ -821,6 +855,9 @@ lemma delete_edges_eq_sdiff_from_edge_set (s : set (sym2 V)) : G.delete_edges s = G \ from_edge_set s := by { ext, exact ⟨λ h, ⟨h.1, not_and_of_not_left _ h.2⟩, λ h, ⟨h.1, not_and'.mp h.2 h.ne⟩⟩ } +@[simp] lemma delete_edges_eq {s : set (sym2 V)} : G.delete_edges s = G ↔ disjoint G.edge_set s := +by rw [delete_edges_eq_sdiff_from_edge_set, sdiff_eq_left, disjoint_from_edge_set] + lemma compl_eq_delete_edges : Gᶜ = (⊤ : simple_graph V).delete_edges G.edge_set := by { ext, simp } @@ -871,18 +908,17 @@ get a graph with the property `p`. -/ def delete_far (p : simple_graph V → Prop) (r : 𝕜) : Prop := ∀ ⦃s⦄, s ⊆ G.edge_finset → p (G.delete_edges s) → r ≤ s.card -open_locale classical - variables {G} lemma delete_far_iff : - G.delete_far p r ↔ ∀ ⦃H⦄, H ≤ G → p H → r ≤ G.edge_finset.card - H.edge_finset.card := + G.delete_far p r ↔ ∀ ⦃H : simple_graph _⦄ [decidable_rel H.adj], + by exactI H ≤ G → p H → r ≤ G.edge_finset.card - H.edge_finset.card := begin - refine ⟨λ h H hHG hH, _, λ h s hs hG, _⟩, + refine ⟨λ h H _ hHG hH, _, λ h s hs hG, _⟩, { have := h (sdiff_subset G.edge_finset H.edge_finset), simp only [delete_edges_sdiff_eq_of_le _ hHG, edge_finset_mono hHG, card_sdiff, card_le_of_subset, coe_sdiff, coe_edge_finset, nat.cast_sub] at this, - exact this hH }, + convert this hH }, { simpa [card_sdiff hs, edge_finset_delete_edges, -set.to_finset_card, nat.cast_sub, card_le_of_subset hs] using h (G.delete_edges_le s) hG } end @@ -906,9 +942,22 @@ protected def map (f : V ↪ W) (G : simple_graph V) : simple_graph W := @[simp] lemma map_adj (f : V ↪ W) (G : simple_graph V) (u v : W) : (G.map f).adj u v ↔ ∃ (u' v' : V), G.adj u' v' ∧ f u' = u ∧ f v' = v := iff.rfl +lemma map_adj_apply {G : simple_graph V} {f : V ↪ W} {a b : V} : + (G.map f).adj (f a) (f b) ↔ G.adj a b := by simp + lemma map_monotone (f : V ↪ W) : monotone (simple_graph.map f) := by { rintros G G' h _ _ ⟨u, v, ha, rfl, rfl⟩, exact ⟨_, _, h ha, rfl, rfl⟩ } +@[simp] lemma map_id : G.map (function.embedding.refl _) = G := +ext _ _ $ relation.map_id_id _ + +@[simp] lemma map_map (f : V ↪ W) (g : W ↪ X) : (G.map f).map g = G.map (f.trans g) := +ext _ _ $ relation.map_map _ _ _ _ _ + +instance decidable_map (f : V ↪ W) (G : simple_graph V) [decidable_rel (relation.map G.adj f f)] : + decidable_rel (G.map f).adj := +‹decidable_rel _› + /-- Given a function, there is a contravariant induced map on graphs by pulling back the adjacency relation. This is one of the ways of creating induced graphs. See `simple_graph.induce` for a wrapper. @@ -917,6 +966,24 @@ This is surjective when `f` is injective (see `simple_graph.comap_surjective`).- @[simps] protected def comap (f : V → W) (G : simple_graph W) : simple_graph V := { adj := λ u v, G.adj (f u) (f v) } +@[simp] lemma comap_id {G : simple_graph V} : G.comap id = G := ext _ _ rfl + +@[simp] lemma comap_comap {G : simple_graph X} (f : V → W) (g : W → X) : + (G.comap g).comap f = G.comap (g ∘ f) := rfl + +instance decidable_comap (f : V → W) (G : simple_graph W) [decidable_rel G.adj] : + decidable_rel (simple_graph.comap f G).adj := +λ _ _, ‹decidable_rel G.adj› _ _ + +lemma comap_symm (G : simple_graph V) (e : V ≃ W) : + G.comap e.symm.to_embedding = G.map e.to_embedding := +by { ext, simp only [equiv.apply_eq_iff_eq_symm_apply, comap_adj, map_adj, equiv.to_embedding_apply, + exists_eq_right_right, exists_eq_right] } + +lemma map_symm (G : simple_graph W) (e : V ≃ W) : + G.map e.symm.to_embedding = G.comap e.to_embedding := +by rw [←comap_symm, e.symm_symm] + lemma comap_monotone (f : V ↪ W) : monotone (simple_graph.comap f) := by { intros G G' h _ _ ha, exact h ha } @@ -939,6 +1006,23 @@ lemma map_le_iff_le_comap (f : V ↪ W) (G : simple_graph V) (G' : simple_graph lemma map_comap_le (f : V ↪ W) (G : simple_graph W) : (G.comap f).map f ≤ G := by { rw map_le_iff_le_comap, exact le_refl _ } +/-- Equivalent types have equivalent simple graphs. -/ +@[simps apply] protected def _root_.equiv.simple_graph (e : V ≃ W) : + simple_graph V ≃ simple_graph W := +{ to_fun := simple_graph.comap e.symm, + inv_fun := simple_graph.comap e, + left_inv := λ _, by simp, + right_inv := λ _, by simp } + +@[simp] lemma _root_.equiv.simple_graph_refl : (equiv.refl V).simple_graph = equiv.refl _ := +by { ext, refl } + +@[simp] lemma _root_.equiv.simple_graph_trans (e₁ : V ≃ W) (e₂ : W ≃ X) : + (e₁.trans e₂).simple_graph = e₁.simple_graph.trans e₂.simple_graph := rfl + +@[simp] lemma _root_.equiv.symm_simple_graph (e : V ≃ W) : + e.simple_graph.symm = e.symm.simple_graph := rfl + /-! ## Induced graphs -/ /- Given a set `s` of vertices, we can restrict a graph to those vertices by restricting its @@ -1296,10 +1380,23 @@ infix ` ↪g ` : 50 := embedding infix ` ≃g ` : 50 := iso namespace hom -variables {G G'} (f : G →g G') +variables {G G'} {H : simple_graph W} (f : G →g G') /-- The identity homomorphism from a graph to itself. -/ -abbreviation id : G →g G := rel_hom.id _ +protected abbreviation id : G →g G := rel_hom.id _ + +@[simp, norm_cast] lemma coe_id : ⇑(hom.id : G →g G) = _root_.id := rfl + +instance [subsingleton (V → W)] : subsingleton (G →g H) := fun_like.coe_injective.subsingleton + +instance [is_empty V] : unique (G →g H) := +{ default := ⟨is_empty_elim, is_empty_elim⟩, + uniq := λ _, subsingleton.elim _ _ } + +noncomputable instance [fintype V] [fintype W] : fintype (G →g H) := +by classical; exact fun_like.fintype _ + +instance [finite V] [finite W] : finite (G →g H) := fun_like.finite _ lemma map_adj {v w : V} (h : G.adj v w) : G'.adj (f v) (f w) := f.map_rel' h @@ -1358,10 +1455,15 @@ abbreviation comp (f' : G' →g G'') (f : G →g G') : G →g G'' := f'.comp f @[simp] lemma coe_comp (f' : G' →g G'') (f : G →g G') : ⇑(f'.comp f) = f' ∘ f := rfl +/-- The graph homomorphism from a smaller graph to a bigger one. -/ +def of_le {H : simple_graph V} (h : G ≤ H) : G →g H := ⟨id, h⟩ + +@[simp, norm_cast] lemma coe_of_le {H : simple_graph V} (h : G ≤ H) : ⇑(of_le h) = id := rfl + end hom namespace embedding -variables {G G'} (f : G ↪g G') +variables {G G'} {H : simple_graph W} (f : G ↪g G') /-- The identity embedding from a graph to itself. -/ abbreviation refl : G ↪g G := rel_embedding.refl _ @@ -1369,7 +1471,9 @@ abbreviation refl : G ↪g G := rel_embedding.refl _ /-- An embedding of graphs gives rise to a homomorphism of graphs. -/ abbreviation to_hom : G →g G' := f.to_rel_hom -lemma map_adj_iff {v w : V} : G'.adj (f v) (f w) ↔ G.adj v w := f.map_rel_iff +@[simp] lemma coe_to_hom (f : G ↪g H) : ⇑f.to_hom = f := rfl + +@[simp] lemma map_adj_iff {v w : V} : G'.adj (f v) (f w) ↔ G.adj v w := f.map_rel_iff lemma map_mem_edge_set_iff {e : sym2 V} : e.map f ∈ G'.edge_set ↔ e ∈ G.edge_set := quotient.ind (λ ⟨v, w⟩, f.map_adj_iff) e diff --git a/src/combinatorics/simple_graph/clique.lean b/src/combinatorics/simple_graph/clique.lean index 996aeb3dd5c72..8d2b10d255ca5 100644 --- a/src/combinatorics/simple_graph/clique.lean +++ b/src/combinatorics/simple_graph/clique.lean @@ -5,6 +5,7 @@ Authors: Yaël Dillies, Bhavik Mehta -/ import combinatorics.simple_graph.basic import data.finset.pairwise +import data.finset.preimage /-! # Graph cliques @@ -25,13 +26,13 @@ adjacent. ## TODO * Clique numbers -* Do we need `clique_set`, a version of `clique_finset` for infinite graphs? +* Dualise all the API to get independent sets -/ -open finset fintype +open finset fintype function namespace simple_graph -variables {α : Type*} (G H : simple_graph α) +variables {α β : Type*} (G H : simple_graph α) /-! ### Cliques -/ @@ -61,13 +62,33 @@ end instance [decidable_eq α] [decidable_rel G.adj] {s : finset α} : decidable (G.is_clique s) := decidable_of_iff' _ G.is_clique_iff -variables {G H} +variables {G H} {a b : α} + +@[simp] lemma is_clique_empty : G.is_clique ∅ := set.pairwise_empty _ +@[simp] lemma is_clique_singleton (a : α) : G.is_clique {a} := set.pairwise_singleton _ _ + +lemma is_clique_pair : G.is_clique {a, b} ↔ a ≠ b → G.adj a b := +set.pairwise_pair_of_symmetric G.symm + +@[simp] lemma is_clique_insert : + G.is_clique (insert a s) ↔ G.is_clique s ∧ ∀ b ∈ s, a ≠ b → G.adj a b := +set.pairwise_insert_of_symmetric G.symm -lemma is_clique.mono (h : G ≤ H) : G.is_clique s → H.is_clique s := -by { simp_rw is_clique_iff, exact set.pairwise.mono' h } +lemma is_clique_insert_of_not_mem (ha : a ∉ s) : + G.is_clique (insert a s) ↔ G.is_clique s ∧ ∀ b ∈ s, G.adj a b := +set.pairwise_insert_of_symmetric_of_not_mem G.symm ha -lemma is_clique.subset (h : t ⊆ s) : G.is_clique s → G.is_clique t := -by { simp_rw is_clique_iff, exact set.pairwise.mono h } +lemma is_clique.insert (hs : G.is_clique s) (h : ∀ b ∈ s, a ≠ b → G.adj a b) : + G.is_clique (insert a s) := +hs.insert_of_symmetric G.symm h + +lemma is_clique.mono (h : G ≤ H) : G.is_clique s → H.is_clique s := set.pairwise.mono' h +lemma is_clique.subset (h : t ⊆ s) : G.is_clique s → G.is_clique t := set.pairwise.mono h + +protected lemma is_clique.map {G : simple_graph α} {s : set α} (h : G.is_clique s) {f : α ↪ β} : + (G.map f).is_clique (f '' s) := +by { rintro _ ⟨a, ha, rfl⟩ _ ⟨b, hb, rfl⟩ hab, + exact ⟨a, b, h ha hb $ ne_of_apply_ne _ hab, rfl, rfl⟩ } @[simp] lemma is_clique_bot_iff : (⊥ : simple_graph α).is_clique s ↔ (s : set α).subsingleton := set.pairwise_bot_iff @@ -93,11 +114,19 @@ instance [decidable_eq α] [decidable_rel G.adj] {n : ℕ} {s : finset α} : decidable (G.is_n_clique n s) := decidable_of_iff' _ G.is_n_clique_iff -variables {G H} +variables {G H} {a b c : α} + +@[simp] lemma is_n_clique_empty : G.is_n_clique n ∅ ↔ n = 0 := by simp [is_n_clique_iff, eq_comm] +@[simp] lemma is_n_clique_singleton : G.is_n_clique n {a} ↔ n = 1 := +by simp [is_n_clique_iff, eq_comm] lemma is_n_clique.mono (h : G ≤ H) : G.is_n_clique n s → H.is_n_clique n s := by { simp_rw is_n_clique_iff, exact and.imp_left (is_clique.mono h) } +protected lemma is_n_clique.map (h : G.is_n_clique n s) {f : α ↪ β} : + (G.map f).is_n_clique n (s.map f) := +⟨by { rw coe_map, exact h.1.map}, (card_map _).trans h.2⟩ + @[simp] lemma is_n_clique_bot_iff : (⊥ : simple_graph α).is_n_clique n s ↔ n ≤ 1 ∧ s.card = n := begin rw [is_n_clique_iff, is_clique_bot_iff], @@ -106,7 +135,22 @@ begin exact card_le_one.symm, end -variables [decidable_eq α] {a b c : α} +@[simp] lemma is_n_clique_zero : G.is_n_clique 0 s ↔ s = ∅ := +by { simp only [is_n_clique_iff, finset.card_eq_zero, and_iff_right_iff_imp], rintro rfl, simp } + +@[simp] lemma is_n_clique_one : G.is_n_clique 1 s ↔ ∃ a, s = {a} := +by { simp only [is_n_clique_iff, card_eq_one, and_iff_right_iff_imp], rintro ⟨a, rfl⟩, simp } + +variables [decidable_eq α] + +lemma is_n_clique.insert (hs : G.is_n_clique n s) (h : ∀ b ∈ s, G.adj a b) : + G.is_n_clique (n + 1) (insert a s) := +begin + split, + { push_cast, + exact hs.1.insert (λ b hb _, h _ hb) }, + { rw [card_insert_of_not_mem (λ ha, (h _ ha).ne rfl), hs.2] } +end lemma is_3_clique_triple_iff : G.is_n_clique 3 {a, b, c} ↔ G.adj a b ∧ G.adj a c ∧ G.adj b c := begin @@ -193,7 +237,7 @@ begin use (iso.complete_graph (fintype.equiv_fin α)).symm.to_embedding.trans f, end -lemma clique_free_bot (h : 2 ≤ n) : (⊥ : simple_graph α).clique_free n := +@[simp] lemma clique_free_bot (h : 2 ≤ n) : (⊥ : simple_graph α).clique_free n := begin rintro t ht, rw is_n_clique_bot_iff at ht, @@ -219,8 +263,77 @@ begin simpa using fintype.card_le_of_embedding h.some.to_embedding, end +@[simp] lemma clique_free_two : G.clique_free 2 ↔ G = ⊥ := +begin + classical, + split, + { simp_rw [←edge_set_eq_empty, set.eq_empty_iff_forall_not_mem, sym2.forall, mem_edge_set], + exact λ h a b hab, h _ ⟨by simpa [hab.ne], card_doubleton hab.ne⟩ }, + { rintro rfl, + exact clique_free_bot le_rfl } +end + end clique_free +section clique_free_on +variables {s s₁ s₂ : set α} {t : finset α} {a b : α} {m n : ℕ} + +/-- `G.clique_free_on s n` means that `G` has no `n`-cliques contained in `s`. -/ +def clique_free_on (G : simple_graph α) (s : set α) (n : ℕ) : Prop := +∀ ⦃t⦄, ↑t ⊆ s → ¬G.is_n_clique n t + +lemma clique_free_on.subset (hs : s₁ ⊆ s₂) (h₂ : G.clique_free_on s₂ n) : G.clique_free_on s₁ n := +λ t hts, h₂ $ hts.trans hs + +lemma clique_free_on.mono (hmn : m ≤ n) (hG : G.clique_free_on s m) : G.clique_free_on s n := +begin + rintro t hts ht, + obtain ⟨u, hut, hu⟩ := t.exists_smaller_set _ (hmn.trans ht.card_eq.ge), + exact hG ((coe_subset.2 hut).trans hts) ⟨ht.clique.subset hut, hu⟩, +end + +lemma clique_free_on.anti (hGH : G ≤ H) (hH : H.clique_free_on s n) : G.clique_free_on s n := +λ t hts ht, hH hts $ ht.mono hGH + +@[simp] lemma clique_free_on_empty : G.clique_free_on ∅ n ↔ n ≠ 0 := +by simp [clique_free_on, set.subset_empty_iff] + +@[simp] lemma clique_free_on_singleton : G.clique_free_on {a} n ↔ 1 < n := +by obtain _ | _ | n := n; simp [clique_free_on, set.subset_singleton_iff_eq] + +@[simp] lemma clique_free_on_univ : G.clique_free_on set.univ n ↔ G.clique_free n := +by simp [clique_free, clique_free_on] + +protected lemma clique_free.clique_free_on (hG : G.clique_free n) : G.clique_free_on s n := +λ t _, hG _ + +lemma clique_free_on_of_card_lt {s : finset α} (h : s.card < n) : G.clique_free_on s n := +λ t hts ht, h.not_le $ ht.2.symm.trans_le $ card_mono hts + +--TOOD: Restate using `simple_graph.indep_set` once we have it +@[simp] lemma clique_free_on_two : G.clique_free_on s 2 ↔ s.pairwise G.adjᶜ := +begin + classical, + refine ⟨λ h a ha b hb _ hab, h _ ⟨by simpa [hab.ne], card_doubleton hab.ne⟩, _⟩, + { push_cast, + exact set.insert_subset.2 ⟨ha, set.singleton_subset_iff.2 hb⟩ }, + simp only [clique_free_on, is_n_clique_iff, card_eq_two, coe_subset, not_and, not_exists], + rintro h t hst ht a b hab rfl, + simp only [coe_insert, coe_singleton, set.insert_subset, set.singleton_subset_iff] at hst, + refine h hst.1 hst.2 hab (ht _ _ hab); simp, +end + +lemma clique_free_on.of_succ (hs : G.clique_free_on s (n + 1)) (ha : a ∈ s) : + G.clique_free_on (s ∩ G.neighbor_set a) n := +begin + classical, + refine λ t hts ht, hs _ (ht.insert $ λ b hb, (hts hb).2), + push_cast, + exact set.insert_subset.2 ⟨ha, hts.trans $ set.inter_subset_left _ _⟩, +end + +end clique_free_on + /-! ### Set of cliques -/ section clique_set @@ -229,22 +342,60 @@ variables (G) {n : ℕ} {a b c : α} {s : finset α} /-- The `n`-cliques in a graph as a set. -/ def clique_set (n : ℕ) : set (finset α) := {s | G.is_n_clique n s} -lemma mem_clique_set_iff : s ∈ G.clique_set n ↔ G.is_n_clique n s := iff.rfl +@[simp] lemma mem_clique_set_iff : s ∈ G.clique_set n ↔ G.is_n_clique n s := iff.rfl @[simp] lemma clique_set_eq_empty_iff : G.clique_set n = ∅ ↔ G.clique_free n := by simp_rw [clique_free, set.eq_empty_iff_forall_not_mem, mem_clique_set_iff] -alias clique_set_eq_empty_iff ↔ _ clique_free.clique_set - -attribute [protected] clique_free.clique_set - variables {G H} +protected lemma clique_free.clique_set : G.clique_free n → G.clique_set n = ∅ := +G.clique_set_eq_empty_iff.2 + @[mono] lemma clique_set_mono (h : G ≤ H) : G.clique_set n ⊆ H.clique_set n := λ _, is_n_clique.mono h lemma clique_set_mono' (h : G ≤ H) : G.clique_set ≤ H.clique_set := λ _, clique_set_mono h +@[simp] lemma clique_set_zero (G : simple_graph α) : G.clique_set 0 = {∅} := +set.ext $ λ s, by simp + +@[simp] lemma clique_set_one (G : simple_graph α) : G.clique_set 1 = set.range singleton := +set.ext $ λ s, by simp [eq_comm] + +@[simp] lemma clique_set_bot (hn : 1 < n) : (⊥ : simple_graph α).clique_set n = ∅ := +(clique_free_bot hn).clique_set + +@[simp] lemma clique_set_map (hn : n ≠ 1) (G : simple_graph α) (f : α ↪ β) : + (G.map f).clique_set n = map f '' G.clique_set n := +begin + ext s, + split, + { rintro ⟨hs, rfl⟩, + have hs' : (s.preimage f $ f.injective.inj_on _).map f = s, + { classical, + rw [map_eq_image, image_preimage, filter_true_of_mem], + rintro a ha, + obtain ⟨b, hb, hba⟩ := exists_mem_ne (hn.lt_of_le' $ finset.card_pos.2 ⟨a, ha⟩) a, + obtain ⟨c, _, _, hc, _⟩ := hs ha hb hba.symm, + exact ⟨c, hc⟩ }, + refine ⟨s.preimage f $ f.injective.inj_on _, ⟨_, by rw [←card_map f, hs']⟩, hs'⟩, + rw coe_preimage, + exact λ a ha b hb hab, map_adj_apply.1 (hs ha hb $ f.injective.ne hab) }, + { rintro ⟨s, hs, rfl⟩, + exact is_n_clique.map hs } +end + +@[simp] lemma clique_set_map_of_equiv (G : simple_graph α) (e : α ≃ β) (n : ℕ) : + (G.map e.to_embedding).clique_set n = map e.to_embedding '' G.clique_set n := +begin + obtain rfl | hn := eq_or_ne n 1, + { ext, + simp [e.exists_congr_left] }, + { exact clique_set_map hn _ _ } +end + + end clique_set /-! ### Finset of cliques -/ @@ -255,10 +406,11 @@ variables (G) [fintype α] [decidable_eq α] [decidable_rel G.adj] {n : ℕ} {a /-- The `n`-cliques in a graph as a finset. -/ def clique_finset (n : ℕ) : finset (finset α) := univ.filter $ G.is_n_clique n -lemma mem_clique_finset_iff : s ∈ G.clique_finset n ↔ G.is_n_clique n s := +@[simp] lemma mem_clique_finset_iff : s ∈ G.clique_finset n ↔ G.is_n_clique n s := mem_filter.trans $ and_iff_right $ mem_univ _ -@[simp] lemma coe_clique_finset (n : ℕ) : (G.clique_finset n : set (finset α)) = G.clique_set n := +@[simp, norm_cast] lemma coe_clique_finset (n : ℕ) : + (G.clique_finset n : set (finset α)) = G.clique_set n := set.ext $ λ _, mem_clique_finset_iff _ @[simp] lemma clique_finset_eq_empty_iff : G.clique_finset n = ∅ ↔ G.clique_free n := @@ -268,10 +420,31 @@ alias clique_finset_eq_empty_iff ↔ _ _root_.simple_graph.clique_free.clique_fi attribute [protected] clique_free.clique_finset -variables {G} [decidable_rel H.adj] +variables {G} + +lemma card_clique_finset_le : (G.clique_finset n).card ≤ (card α).choose n := +begin + rw [←card_univ, ←card_powerset_len], + refine card_mono (λ s, _), + simpa [mem_powerset_len_univ_iff] using is_n_clique.card_eq, +end + +variables [decidable_rel H.adj] @[mono] lemma clique_finset_mono (h : G ≤ H) : G.clique_finset n ⊆ H.clique_finset n := monotone_filter_right _ $ λ _, is_n_clique.mono h +variables [fintype β] [decidable_eq β] (G) + +@[simp] lemma clique_finset_map (f : α ↪ β) (hn : n ≠ 1) : + (G.map f).clique_finset n = (G.clique_finset n).map ⟨map f, finset.map_injective _⟩ := +coe_injective $ + by simp_rw [coe_clique_finset, clique_set_map hn, coe_map, coe_clique_finset, embedding.coe_fn_mk] + +@[simp] lemma clique_finset_map_of_equiv (e : α ≃ β) (n : ℕ) : + (G.map e.to_embedding).clique_finset n = + (G.clique_finset n).map ⟨map e.to_embedding, finset.map_injective _⟩ := +coe_injective $ by push_cast; exact clique_set_map_of_equiv _ _ _ + end clique_finset end simple_graph diff --git a/src/combinatorics/simple_graph/triangle/basic.lean b/src/combinatorics/simple_graph/triangle/basic.lean index 639207f75591c..38c350e394e93 100644 --- a/src/combinatorics/simple_graph/triangle/basic.lean +++ b/src/combinatorics/simple_graph/triangle/basic.lean @@ -42,7 +42,8 @@ G.delete_far (λ H, H.clique_free 3) $ ε * (card α^2 : ℕ) lemma far_from_triangle_free_iff : G.far_from_triangle_free ε ↔ - ∀ ⦃H⦄, H ≤ G → H.clique_free 3 → ε * (card α^2 : ℕ) ≤ G.edge_finset.card - H.edge_finset.card := + ∀ ⦃H : simple_graph _⦄ [decidable_rel H.adj], by exactI + H ≤ G → H.clique_free 3 → ε * (card α^2 : ℕ) ≤ G.edge_finset.card - H.edge_finset.card := delete_far_iff alias far_from_triangle_free_iff ↔ far_from_triangle_free.le_card_sub_card _ diff --git a/src/data/finset/card.lean b/src/data/finset/card.lean index 98cf0399b0e58..f023ef54c1104 100644 --- a/src/data/finset/card.lean +++ b/src/data/finset/card.lean @@ -434,6 +434,13 @@ begin exact card_le_of_subset hx } end +lemma exists_mem_ne (hs : 1 < s.card) (a : α) : ∃ b ∈ s, b ≠ a := +begin + by_contra', + haveI : nonempty α := ⟨a⟩, + exact hs.not_le (card_le_one_iff_subset_singleton.2 ⟨a, subset_singleton_iff'.2 this⟩), +end + /-- A `finset` of a subsingleton type has cardinality at most one. -/ lemma card_le_one_of_subsingleton [subsingleton α] (s : finset α) : s.card ≤ 1 := finset.card_le_one_iff.2 $ λ _ _ _ _, subsingleton.elim _ _ diff --git a/src/data/finset/preimage.lean b/src/data/finset/preimage.lean index ba9639d06700d..afe69fac91482 100644 --- a/src/data/finset/preimage.lean +++ b/src/data/finset/preimage.lean @@ -61,6 +61,10 @@ finset.coe_injective (by simp) preimage sᶜ f (hf.inj_on _) = (preimage s f (hf.inj_on _))ᶜ := finset.coe_injective (by simp) +@[simp] lemma preimage_map (f : α ↪ β) (s : finset α) : + (s.map f).preimage f (f.injective.inj_on _) = s := +coe_injective $ by simp only [coe_preimage, coe_map, set.preimage_image_eq _ f.injective] + lemma monotone_preimage {f : α → β} (h : injective f) : monotone (λ s, preimage s f (h.inj_on _)) := λ s t hst x hx, mem_preimage.2 (hst $ mem_preimage.1 hx) diff --git a/src/logic/basic.lean b/src/logic/basic.lean index 71baa00527e9d..26ef5b00ad237 100644 --- a/src/logic/basic.lean +++ b/src/logic/basic.lean @@ -423,6 +423,15 @@ lemma iff.not (h : a ↔ b) : ¬ a ↔ ¬ b := not_congr h lemma iff.not_left (h : a ↔ ¬ b) : ¬ a ↔ b := h.not.trans not_not lemma iff.not_right (h : ¬ a ↔ b) : a ↔ ¬ b := not_not.symm.trans h.not +protected lemma iff.ne {α β : Sort*} {a b : α} {c d : β} : (a = b ↔ c = d) → (a ≠ b ↔ c ≠ d) := +iff.not + +lemma iff.ne_left {α β : Sort*} {a b : α} {c d : β} : (a = b ↔ c ≠ d) → (a ≠ b ↔ c = d) := +iff.not_left + +lemma iff.ne_right {α β : Sort*} {a b : α} {c d : β} : (a ≠ b ↔ c = d) → (a = b ↔ c ≠ d) := +iff.not_right + /-! ### Declarations about `xor` -/ @[simp] theorem xor_true : xor true = not := funext $ λ a, by simp [xor] diff --git a/src/logic/relation.lean b/src/logic/relation.lean index 8467caa176c36..466ad575131e0 100644 --- a/src/logic/relation.lean +++ b/src/logic/relation.lean @@ -42,7 +42,7 @@ the bundled version, see `rel`. open function -variables {α β γ δ : Type*} +variables {α β γ δ ε κ : Type*} section ne_imp @@ -186,6 +186,27 @@ related by `r`. protected def map (r : α → β → Prop) (f : α → γ) (g : β → δ) : γ → δ → Prop := λ c d, ∃ a b, r a b ∧ f a = c ∧ g b = d +section map +variables {r : α → β → Prop} {f : α → γ} {g : β → δ} {c : γ} {d : δ} + +lemma map_apply : relation.map r f g c d ↔ ∃ a b, r a b ∧ f a = c ∧ g b = d := iff.rfl + +@[simp] lemma map_id_id (r : α → β → Prop) : relation.map r id id = r := by simp [relation.map] + +@[simp] lemma map_map (r : α → β → Prop) (f₁ : α → γ) (g₁ : β → δ) (f₂ : γ → ε) (g₂ : δ → κ) : + relation.map (relation.map r f₁ g₁) f₂ g₂ = relation.map r (f₂ ∘ f₁) (g₂ ∘ g₁) := +begin + ext a b, + simp only [map_apply, function.comp_app, ←exists_and_distrib_right, @exists₂_comm γ], + refine exists₂_congr (λ a b, _), + simp [and_assoc], +end + +instance [decidable (∃ a b, r a b ∧ f a = c ∧ g b = d)] : decidable (relation.map r f g c d) := +‹decidable _› + +end map + variables {r : α → α → Prop} {a b c d : α} /-- `refl_trans_gen r`: reflexive transitive closure of `r` -/ From 65a1391a0106c9204fe45bc73a039f056558cb83 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 30 Oct 2023 15:07:49 +0000 Subject: [PATCH 1291/1291] feat(data/{list,multiset,finset}/*): `attach` and `filter` lemmas (#18087) Left commutativity and cardinality of `list.filter`/`multiset.filter`/`finset.filter`. Interaction of `count`/`countp` and `attach`. --- src/algebra/big_operators/basic.lean | 24 ++++++++----- src/algebra/big_operators/order.lean | 2 +- src/data/finset/card.lean | 1 + src/data/finset/image.lean | 16 +++++++++ src/data/list/basic.lean | 25 +++++++++++++ src/data/list/count.lean | 6 ++++ src/data/list/perm.lean | 2 +- src/data/multiset/basic.lean | 50 ++++++++++++++++++++------ src/data/multiset/lattice.lean | 2 +- src/data/set/finite.lean | 3 +- src/number_theory/kummer_dedekind.lean | 2 +- 11 files changed, 108 insertions(+), 25 deletions(-) diff --git a/src/algebra/big_operators/basic.lean b/src/algebra/big_operators/basic.lean index 85ce482038a50..ac6250a471925 100644 --- a/src/algebra/big_operators/basic.lean +++ b/src/algebra/big_operators/basic.lean @@ -106,6 +106,9 @@ variables {s s₁ s₂ : finset α} {a : α} {f g : α → β} @[to_additive] lemma prod_eq_multiset_prod [comm_monoid β] (s : finset α) (f : α → β) : ∏ x in s, f x = (s.1.map f).prod := rfl +@[simp, to_additive] lemma prod_map_val [comm_monoid β] (s : finset α) (f : α → β) : + (s.1.map f).prod = ∏ a in s, f a := rfl + @[to_additive] theorem prod_eq_fold [comm_monoid β] (s : finset α) (f : α → β) : ∏ x in s, f x = s.fold (*) 1 f := @@ -1418,16 +1421,19 @@ begin apply sum_congr rfl h₁ end +lemma nat_cast_card_filter [add_comm_monoid_with_one β] (p) [decidable_pred p] (s : finset α) : + ((filter p s).card : β) = ∑ a in s, if p a then 1 else 0 := +by simp only [add_zero, sum_const, nsmul_eq_mul, eq_self_iff_true, sum_const_zero, sum_ite, + nsmul_one] + +lemma card_filter (p) [decidable_pred p] (s : finset α) : + (filter p s).card = ∑ a in s, ite (p a) 1 0 := +nat_cast_card_filter _ _ + @[simp] -lemma sum_boole {s : finset α} {p : α → Prop} [non_assoc_semiring β] {hp : decidable_pred p} : - (∑ x in s, if p x then (1 : β) else (0 : β)) = (s.filter p).card := -by simp only [add_zero, - mul_one, - finset.sum_const, - nsmul_eq_mul, - eq_self_iff_true, - finset.sum_const_zero, - finset.sum_ite] +lemma sum_boole {s : finset α} {p : α → Prop} [add_comm_monoid_with_one β] {hp : decidable_pred p} : + (∑ x in s, if p x then 1 else 0 : β) = (s.filter p).card := +(nat_cast_card_filter _ _).symm lemma _root_.commute.sum_right [non_unital_non_assoc_semiring β] (s : finset α) (f : α → β) (b : β) (h : ∀ i ∈ s, commute b (f i)) : diff --git a/src/algebra/big_operators/order.lean b/src/algebra/big_operators/order.lean index 2573be61bcdb6..cda6b956aaa63 100644 --- a/src/algebra/big_operators/order.lean +++ b/src/algebra/big_operators/order.lean @@ -174,7 +174,7 @@ lemma prod_le_pow_card (s : finset ι) (f : ι → N) (n : N) (h : ∀ x ∈ s, begin refine (multiset.prod_le_pow_card (s.val.map f) n _).trans _, { simpa using h }, - { simpa } + { simp } end @[to_additive card_nsmul_le_sum] diff --git a/src/data/finset/card.lean b/src/data/finset/card.lean index f023ef54c1104..ae354aa222531 100644 --- a/src/data/finset/card.lean +++ b/src/data/finset/card.lean @@ -42,6 +42,7 @@ variables {s t : finset α} {a b : α} def card (s : finset α) : ℕ := s.1.card lemma card_def (s : finset α) : s.card = s.1.card := rfl +@[simp] lemma card_val (s : finset α) : s.1.card = s.card := rfl @[simp] lemma card_mk {m nodup} : (⟨m, nodup⟩ : finset α).card = m.card := rfl diff --git a/src/data/finset/image.lean b/src/data/finset/image.lean index e41fa0925240e..74a678a27cabc 100644 --- a/src/data/finset/image.lean +++ b/src/data/finset/image.lean @@ -127,6 +127,22 @@ lemma filter_map {p : β → Prop} [decidable_pred p] : (s.map f).filter p = (s.filter (p ∘ f)).map f := eq_of_veq (map_filter _ _ _) +lemma map_filter' (p : α → Prop) [decidable_pred p] (f : α ↪ β) (s : finset α) + [decidable_pred (λ b, ∃ a, p a ∧ f a = b)] : + (s.filter p).map f = (s.map f).filter (λ b, ∃ a, p a ∧ f a = b) := +by simp [(∘), filter_map, f.injective.eq_iff] + +lemma filter_attach' [decidable_eq α] (s : finset α) (p : s → Prop) [decidable_pred p] : + s.attach.filter p = + (s.filter $ λ x, ∃ h, p ⟨x, h⟩).attach.map ⟨subtype.map id $ filter_subset _ _, + subtype.map_injective _ injective_id⟩ := +eq_of_veq $ multiset.filter_attach' _ _ + +@[simp] lemma filter_attach (p : α → Prop) [decidable_pred p] (s : finset α) : + (s.attach.filter (λ x, p ↑x)) = + (s.filter p).attach.map ((embedding.refl _).subtype_map mem_of_mem_filter) := +eq_of_veq $ multiset.filter_attach _ _ + lemma map_filter {f : α ≃ β} {p : α → Prop} [decidable_pred p] : (s.filter p).map f.to_embedding = (s.map f.to_embedding).filter (p ∘ f.symm) := by simp only [filter_map, function.comp, equiv.to_embedding_apply, equiv.symm_apply_apply] diff --git a/src/data/list/basic.lean b/src/data/list/basic.lean index 939c431e6d038..ce40ca44d44af 100644 --- a/src/data/list/basic.lean +++ b/src/data/list/basic.lean @@ -2716,6 +2716,8 @@ end split_at_on with the same elements but in the type `{x // x ∈ l}`. -/ def attach (l : list α) : list {x // x ∈ l} := pmap subtype.mk l (λ a, id) +@[simp] lemma attach_nil : ([] : list α).attach = [] := rfl + theorem sizeof_lt_sizeof_of_mem [has_sizeof α] {x : α} {l : list α} (hx : x ∈ l) : sizeof x < sizeof l := begin @@ -3250,12 +3252,35 @@ theorem map_filter (f : β → α) (l : list β) : filter p (map f l) = map f (filter (p ∘ f) l) := by rw [← filter_map_eq_map, filter_filter_map, filter_map_filter]; refl +lemma map_filter' {f : α → β} (hf : injective f) (l : list α) + [decidable_pred (λ b, ∃ a, p a ∧ f a = b)] : + (l.filter p).map f = (l.map f).filter (λ b, ∃ a, p a ∧ f a = b) := +by simp [(∘), map_filter, hf.eq_iff] + +lemma filter_attach' (l : list α) (p : {a // a ∈ l} → Prop) [decidable_eq α] [decidable_pred p] : + l.attach.filter p = (l.filter $ λ x, ∃ h, p ⟨x, h⟩).attach.map + (subtype.map id $ λ x hx, let ⟨h, _⟩ := of_mem_filter hx in h) := +begin + classical, + refine map_injective_iff.2 subtype.coe_injective _, + simp [(∘), map_filter' _ subtype.coe_injective], +end + +@[simp] lemma filter_attach (l : list α) (p : α → Prop) [decidable_pred p] : + l.attach.filter (λ x, p ↑x) = (l.filter p).attach.map (subtype.map id $ λ _, mem_of_mem_filter) := +map_injective_iff.2 subtype.coe_injective $ by + simp_rw [map_map, (∘), subtype.map, subtype.coe_mk, id.def, ←map_filter, attach_map_coe] + @[simp] theorem filter_filter (q) [decidable_pred q] : ∀ l, filter p (filter q l) = filter (λ a, p a ∧ q a) l | [] := rfl | (a :: l) := by by_cases hp : p a; by_cases hq : q a; simp only [hp, hq, filter, if_true, if_false, true_and, false_and, filter_filter l, eq_self_iff_true] +lemma filter_comm (q) [decidable_pred q] (l : list α) : + filter p (filter q l) = filter q (filter p l) := +by simp [and_comm] + @[simp] lemma filter_true {h : decidable_pred (λ a : α, true)} (l : list α) : @filter α (λ _, true) h l = l := by convert filter_eq_self.2 (λ _ _, trivial) diff --git a/src/data/list/count.lean b/src/data/list/count.lean index 5660878757f14..25d72187efa05 100644 --- a/src/data/list/count.lean +++ b/src/data/list/count.lean @@ -84,6 +84,9 @@ by simp only [countp_eq_length_filter, filter_filter] | [] := rfl | (a::l) := by rw [map_cons, countp_cons, countp_cons, countp_map] +@[simp] lemma countp_attach (l : list α) : l.attach.countp (λ a, p ↑a) = l.countp p := +by rw [←countp_map, attach_map_coe] + variables {p q} lemma countp_mono_left (h : ∀ x ∈ l, p x → q x) : countp p l ≤ countp q l := @@ -197,6 +200,9 @@ lemma count_bind {α β} [decidable_eq β] (l : list α) (f : α → list β) (x count x (l.bind f) = sum (map (count x ∘ f) l) := by rw [list.bind, count_join, map_map] +@[simp] lemma count_attach (a : {x // x ∈ l}) : l.attach.count a = l.count a := +eq.trans (countp_congr $ λ _ _, subtype.ext_iff) $ countp_attach _ _ + @[simp] lemma count_map_of_injective {α β} [decidable_eq α] [decidable_eq β] (l : list α) (f : α → β) (hf : function.injective f) (x : α) : count (f x) (map f l) = count x l := diff --git a/src/data/list/perm.lean b/src/data/list/perm.lean index 8c1667485d02c..22c3396410970 100644 --- a/src/data/list/perm.lean +++ b/src/data/list/perm.lean @@ -396,7 +396,7 @@ theorem subperm.countp_le (p : α → Prop) [decidable_pred p] | ⟨l, p', s⟩ := p'.countp_eq p ▸ s.countp_le p theorem perm.countp_congr (s : l₁ ~ l₂) {p p' : α → Prop} [decidable_pred p] [decidable_pred p'] - (hp : ∀ x ∈ l₁, p x = p' x) : l₁.countp p = l₂.countp p' := + (hp : ∀ x ∈ l₁, p x ↔ p' x) : l₁.countp p = l₂.countp p' := begin rw ← s.countp_eq p', clear s, diff --git a/src/data/multiset/basic.lean b/src/data/multiset/basic.lean index f58ff719fad96..99873f8d0abe7 100644 --- a/src/data/multiset/basic.lean +++ b/src/data/multiset/basic.lean @@ -16,7 +16,7 @@ These are implemented as the quotient of a list by permutations. We define the global infix notation `::ₘ` for `multiset.cons`. -/ -open list subtype nat +open function list nat subtype variables {α : Type*} {β : Type*} {γ : Type*} @@ -1543,6 +1543,10 @@ le_antisymm (le_inter filter p (filter q s) = filter (λ a, p a ∧ q a) s := quot.induction_on s $ λ l, congr_arg coe $ filter_filter p q l +lemma filter_comm (q) [decidable_pred q] (s : multiset α) : + filter p (filter q s) = filter q (filter p s) := +by simp [and_comm] + theorem filter_add_filter (q) [decidable_pred q] (s : multiset α) : filter p s + filter q s = filter (λ a, p a ∨ q a) s + filter (λ a, p a ∧ q a) s := multiset.induction_on s rfl $ λ a s IH, @@ -1556,6 +1560,11 @@ theorem map_filter (f : β → α) (s : multiset β) : filter p (map f s) = map f (filter (p ∘ f) s) := quot.induction_on s (λ l, by simp [map_filter]) +lemma map_filter' {f : α → β} (hf : injective f) (s : multiset α) + [decidable_pred (λ b, ∃ a, p a ∧ f a = b)] : + (s.filter p).map f = (s.map f).filter (λ b, ∃ a, p a ∧ f a = b) := +by simp [(∘), map_filter, hf.eq_iff] + /-! ### Simultaneously filter and map elements of a multiset -/ /-- `filter_map f s` is a combination filter/map operation on `s`. @@ -1704,6 +1713,18 @@ begin card_singleton, add_comm] }, end +@[simp] lemma countp_attach (s : multiset α) : s.attach.countp (λ a, p ↑a) = s.countp p := +quotient.induction_on s $ λ l, begin + simp only [quot_mk_to_coe, coe_countp], + rw [quot_mk_to_coe, coe_attach, coe_countp], + exact list.countp_attach _ _, +end + +@[simp] lemma filter_attach (m : multiset α) (p : α → Prop) [decidable_pred p] : + (m.attach.filter (λ x, p ↑x)) = + (m.filter p).attach.map (subtype.map id $ λ _, multiset.mem_of_mem_filter) := +quotient.induction_on m $ λ l, congr_arg coe (list.filter_attach l p) + variable {p} theorem countp_pos {s} : 0 < countp p s ↔ ∃ a ∈ s, p a := @@ -1720,7 +1741,7 @@ countp_pos.2 ⟨_, h, pa⟩ theorem countp_congr {s s' : multiset α} (hs : s = s') {p p' : α → Prop} [decidable_pred p] [decidable_pred p'] - (hp : ∀ x ∈ s, p x = p' x) : s.countp p = s'.countp p' := + (hp : ∀ x ∈ s, p x ↔ p' x) : s.countp p = s'.countp p' := quot.induction_on₂ s s' (λ l l' hs hp, begin simp only [quot_mk_to_coe'', coe_eq_coe] at hs, exact hs.countp_congr hp, @@ -1731,7 +1752,7 @@ end /-! ### Multiplicity of an element -/ section -variable [decidable_eq α] +variables [decidable_eq α] {s : multiset α} /-- `count a s` is the multiplicity of `a` in `s`. -/ def count (a : α) : multiset α → ℕ := countp (eq a) @@ -1778,6 +1799,9 @@ def count_add_monoid_hom (a : α) : multiset α →+ ℕ := countp_add_monoid_ho @[simp] theorem count_nsmul (a : α) (n s) : count a (n • s) = n * count a s := by induction n; simp [*, succ_nsmul', succ_mul, zero_nsmul] +@[simp] lemma count_attach (a : {x // x ∈ s}) : s.attach.count a = s.count a := +eq.trans (countp_congr rfl $ λ _ _, subtype.ext_iff) $ countp_attach _ _ + theorem count_pos {a : α} {s : multiset α} : 0 < count a s ↔ a ∈ s := by simp [count, countp_pos] @@ -1901,13 +1925,6 @@ begin contradiction } end -@[simp] -lemma attach_count_eq_count_coe (m : multiset α) (a) : m.attach.count a = m.count (a : α) := -calc m.attach.count a - = (m.attach.map (coe : _ → α)).count (a : α) : - (multiset.count_map_eq_count' _ _ subtype.coe_injective _).symm -... = m.count (a : α) : congr_arg _ m.attach_map_coe - lemma filter_eq' (s : multiset α) (b : α) : s.filter (= b) = replicate (count b s) b := quotient.induction_on s $ λ l, congr_arg coe $ filter_eq' l b @@ -2174,6 +2191,19 @@ theorem map_injective {f : α → β} (hf : function.injective f) : function.injective (multiset.map f) := assume x y, (map_eq_map hf).1 +lemma filter_attach' (s : multiset α) (p : {a // a ∈ s} → Prop) [decidable_eq α] + [decidable_pred p] : + s.attach.filter p = + (s.filter $ λ x, ∃ h, p ⟨x, h⟩).attach.map (subtype.map id $ λ x hx, + let ⟨h, _⟩ := of_mem_filter hx in h) := +begin + classical, + refine multiset.map_injective subtype.coe_injective _, + simp only [function.comp, map_filter' _ subtype.coe_injective, subtype.exists, coe_mk, + exists_and_distrib_right, exists_eq_right, attach_map_coe, map_map, map_coe, id.def], + rw attach_map_coe, +end + end map section quot diff --git a/src/data/multiset/lattice.lean b/src/data/multiset/lattice.lean index 7b279882080d6..5e575d279b753 100644 --- a/src/data/multiset/lattice.lean +++ b/src/data/multiset/lattice.lean @@ -39,7 +39,7 @@ sup_bot_eq @[simp] lemma sup_add (s₁ s₂ : multiset α) : (s₁ + s₂).sup = s₁.sup ⊔ s₂.sup := eq.trans (by simp [sup]) (fold_add _ _ _ _ _) -lemma sup_le {s : multiset α} {a : α} : s.sup ≤ a ↔ (∀b ∈ s, b ≤ a) := +@[simp] lemma sup_le {s : multiset α} {a : α} : s.sup ≤ a ↔ (∀b ∈ s, b ≤ a) := multiset.induction_on s (by simp) (by simp [or_imp_distrib, forall_and_distrib] {contextual := tt}) diff --git a/src/data/set/finite.lean b/src/data/set/finite.lean index f86cca6bfa443..c073dce02f496 100644 --- a/src/data/set/finite.lean +++ b/src/data/set/finite.lean @@ -927,8 +927,7 @@ eq.trans (by congr) empty_card theorem card_fintype_insert_of_not_mem {a : α} (s : set α) [fintype s] (h : a ∉ s) : @fintype.card _ (fintype_insert_of_not_mem s h) = fintype.card s + 1 := -by rw [fintype_insert_of_not_mem, fintype.card_of_finset]; - simp [finset.card, to_finset]; refl +by simp [fintype_insert_of_not_mem, fintype.card_of_finset] @[simp] theorem card_insert {a : α} (s : set α) [fintype s] (h : a ∉ s) {d : fintype.{u} (insert a s : set α)} : diff --git a/src/number_theory/kummer_dedekind.lean b/src/number_theory/kummer_dedekind.lean index 28af342219f88..381da69360e7e 100644 --- a/src/number_theory/kummer_dedekind.lean +++ b/src/number_theory/kummer_dedekind.lean @@ -294,7 +294,7 @@ begin rw [multiset.count_map_eq_count' (λ f, ((normalized_factors_map_equiv_normalized_factors_min_poly_mk hI hI' hx hx').symm f : ideal S)), - multiset.attach_count_eq_count_coe], + multiset.count_attach], { exact subtype.coe_injective.comp (equiv.injective _) }, { exact (normalized_factors_map_equiv_normalized_factors_min_poly_mk hI hI' hx hx' _).prop}, { exact irreducible_of_normalized_factor _